diff --git a/Archive/Imo/Imo1972Q5.lean b/Archive/Imo/Imo1972Q5.lean index 3fefb3563e990..9b016c1f1d2f7 100644 --- a/Archive/Imo/Imo1972Q5.lean +++ b/Archive/Imo/Imo1972Q5.lean @@ -27,7 +27,7 @@ from `hneg` directly), finally raising a contradiction with `k' < k'`. theorem imo1972_q5 (f g : ℝ → ℝ) (hf1 : ∀ x, ∀ y, f (x + y) + f (x - y) = 2 * f x * g y) (hf2 : ∀ y, ‖f y‖ ≤ 1) (hf3 : ∃ x, f x ≠ 0) (y : ℝ) : ‖g y‖ ≤ 1 := by -- Suppose the conclusion does not hold. - by_contra' hneg + by_contra! hneg set S := Set.range fun x => ‖f x‖ -- Introduce `k`, the supremum of `f`. let k : ℝ := sSup S @@ -82,8 +82,8 @@ This is a more concise version of the proof proposed by Ruben Van de Velde. -/ theorem imo1972_q5' (f g : ℝ → ℝ) (hf1 : ∀ x, ∀ y, f (x + y) + f (x - y) = 2 * f x * g y) (hf2 : BddAbove (Set.range fun x => ‖f x‖)) (hf3 : ∃ x, f x ≠ 0) (y : ℝ) : ‖g y‖ ≤ 1 := by - -- porting note: moved `by_contra'` up to avoid a bug - by_contra' H + -- porting note: moved `by_contra!` up to avoid a bug + by_contra! H obtain ⟨x, hx⟩ := hf3 set k := ⨆ x, ‖f x‖ have h : ∀ x, ‖f x‖ ≤ k := le_ciSup hf2 diff --git a/Archive/Imo/Imo1994Q1.lean b/Archive/Imo/Imo1994Q1.lean index 05c0bdfec40bc..bc6aa2a85adca 100644 --- a/Archive/Imo/Imo1994Q1.lean +++ b/Archive/Imo/Imo1994Q1.lean @@ -75,7 +75,7 @@ theorem imo1994_q1 (n : ℕ) (m : ℕ) (A : Finset ℕ) (hm : A.card = m + 1) _ = (m + 1) * (n + 1) := by rw [sum_const, card_fin, Nat.nsmul_eq_mul] -- It remains to prove the key inequality, by contradiction rintro k - - by_contra' h : a k + a (rev k) < n + 1 + by_contra! h : a k + a (rev k) < n + 1 -- We exhibit `k+1` elements of `A` greater than `a (rev k)` set f : Fin (m + 1) ↪ ℕ := ⟨fun i => a i + a (rev k), by diff --git a/Archive/Imo/Imo2008Q4.lean b/Archive/Imo/Imo2008Q4.lean index 0ea6215fb9966..4619f634d21ab 100644 --- a/Archive/Imo/Imo2008Q4.lean +++ b/Archive/Imo/Imo2008Q4.lean @@ -68,7 +68,7 @@ theorem imo2008_q4 (f : ℝ → ℝ) (H₁ : ∀ x > 0, f x > 0) : field_simp at H₂ ⊢ linear_combination 1 / 2 * H₂ have h₃ : ∀ x > 0, f x = x ∨ f x = 1 / x := by simpa [sub_eq_zero] using h₂ - by_contra' h + by_contra! h rcases h with ⟨⟨b, hb, hfb₁⟩, ⟨a, ha, hfa₁⟩⟩ obtain hfa₂ := Or.resolve_right (h₃ a ha) hfa₁ -- f(a) ≠ 1/a, f(a) = a diff --git a/Archive/Imo/Imo2013Q5.lean b/Archive/Imo/Imo2013Q5.lean index 4b06d4fc31212..e88f9a448ef7b 100644 --- a/Archive/Imo/Imo2013Q5.lean +++ b/Archive/Imo/Imo2013Q5.lean @@ -35,7 +35,7 @@ namespace Imo2013Q5 theorem le_of_all_pow_lt_succ {x y : ℝ} (hx : 1 < x) (hy : 1 < y) (h : ∀ n : ℕ, 0 < n → x ^ n - 1 < y ^ n) : x ≤ y := by - by_contra' hxy + by_contra! hxy have hxmy : 0 < x - y := sub_pos.mpr hxy have hn : ∀ n : ℕ, 0 < n → (x - y) * (n : ℝ) ≤ x ^ n - y ^ n := by intro n _ @@ -68,7 +68,7 @@ theorem le_of_all_pow_lt_succ {x y : ℝ} (hx : 1 < x) (hy : 1 < y) theorem le_of_all_pow_lt_succ' {x y : ℝ} (hx : 1 < x) (hy : 0 < y) (h : ∀ n : ℕ, 0 < n → x ^ n - 1 < y ^ n) : x ≤ y := by refine' le_of_all_pow_lt_succ hx _ h - by_contra' hy'' : y ≤ 1 + by_contra! hy'' : y ≤ 1 -- Then there exists y' such that 0 < y ≤ 1 < y' < x. let y' := (x + 1) / 2 have h_y'_lt_x : y' < x := @@ -163,7 +163,7 @@ theorem fixed_point_of_gt_1 {f : ℚ → ℝ} {x : ℚ} (hx : 1 < x) (x : ℝ) + (a ^ N - x : ℚ) ≤ f x + (a ^ N - x : ℚ) := by gcongr; exact H5 x hx _ ≤ f x + f (a ^ N - x) := by gcongr; exact H5 _ h_big_enough have hxp : 0 < x := by positivity - have hNp : 0 < N := by by_contra' H; rw [le_zero_iff.mp H] at hN; linarith + have hNp : 0 < N := by by_contra! H; rw [le_zero_iff.mp H] at hN; linarith have h2 := calc f x + f (a ^ N - x) ≤ f (x + (a ^ N - x)) := H2 x (a ^ N - x) hxp (by positivity) diff --git a/Archive/Imo/Imo2021Q1.lean b/Archive/Imo/Imo2021Q1.lean index bfb4538cec9ff..5ed4d28285d37 100644 --- a/Archive/Imo/Imo2021Q1.lean +++ b/Archive/Imo/Imo2021Q1.lean @@ -124,7 +124,7 @@ theorem imo2021_q1 : have hBsub : B ⊆ Finset.Icc n (2 * n) := by intro c hcB; simpa only [Finset.mem_Icc] using h₂ c hcB have hB' : 2 * 1 < (B ∩ (Finset.Icc n (2 * n) \ A) ∪ B ∩ A).card := by - rw [←inter_distrib_left, sdiff_union_self_eq_union, union_eq_left.2 hA, + rw [← inter_distrib_left, sdiff_union_self_eq_union, union_eq_left.2 hA, inter_eq_left.2 hBsub] exact Nat.succ_le_iff.mp hB -- Since B has cardinality greater or equal to 3, there must exist a subset C ⊆ B such that diff --git a/Archive/Wiedijk100Theorems/AscendingDescendingSequences.lean b/Archive/Wiedijk100Theorems/AscendingDescendingSequences.lean index 1ec32b3a3bb3a..3c88750dac1e8 100644 --- a/Archive/Wiedijk100Theorems/AscendingDescendingSequences.lean +++ b/Archive/Wiedijk100Theorems/AscendingDescendingSequences.lean @@ -136,8 +136,7 @@ theorem erdos_szekeres {r s n : ℕ} {f : Fin n → α} (hn : r * s < n) (hf : I -- Finished both goals! -- Now that we have uniqueness of each label, it remains to do some counting to finish off. -- Suppose all the labels are small. - by_contra q - push_neg at q + by_contra! q -- Then the labels `(a_i, b_i)` all fit in the following set: `{ (x,y) | 1 ≤ x ≤ r, 1 ≤ y ≤ s }` let ran : Finset (ℕ × ℕ) := (range r).image Nat.succ ×ˢ (range s).image Nat.succ -- which we prove here. diff --git a/Cache/Hashing.lean b/Cache/Hashing.lean index 358835226ffe2..a4440a6817ce1 100644 --- a/Cache/Hashing.lean +++ b/Cache/Hashing.lean @@ -76,7 +76,7 @@ def getRootHash : IO UInt64 := do pure ((← mathlibDepPath) / ·) let hashs ← rootFiles.mapM fun path => hashFileContents <$> IO.FS.readFile (qualifyPath path) - return hash ((hash Lean.versionString) :: hashs) + return hash (hash Lean.githash :: hashs) /-- Computes the hash of a file, which mixes: diff --git a/Counterexamples/CharPZeroNeCharZero.lean b/Counterexamples/CharPZeroNeCharZero.lean index f3a5e8e24f571..b651338409689 100644 --- a/Counterexamples/CharPZeroNeCharZero.lean +++ b/Counterexamples/CharPZeroNeCharZero.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Damiano Testa, Eric Wieser -/ import Mathlib.Algebra.CharP.Basic +import Mathlib.Algebra.PUnitInstances #align_import char_p_zero_ne_char_zero from "leanprover-community/mathlib"@"328375597f2c0dd00522d9c2e5a33b6a6128feeb" diff --git a/Counterexamples/Phillips.lean b/Counterexamples/Phillips.lean index 45437cd0ecd00..7470ff0135a33 100644 --- a/Counterexamples/Phillips.lean +++ b/Counterexamples/Phillips.lean @@ -261,7 +261,7 @@ theorem exists_discrete_support_nonpos (f : BoundedAdditiveMeasure α) : In this proof, we use explicit coercions `↑s` for `s : A` as otherwise the system tries to find a `CoeFun` instance on `↥A`, which is too costly. -/ - by_contra' h + by_contra! h -- We will formulate things in terms of the type of countable subsets of `α`, as this is more -- convenient to formalize the inductive construction. let A : Set (Set α) := {t | t.Countable} diff --git a/Mathlib.lean b/Mathlib.lean index 89703b8ced284..40f868ae0441b 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -104,6 +104,7 @@ import Mathlib.Algebra.CharP.LocalRing import Mathlib.Algebra.CharP.MixedCharZero import Mathlib.Algebra.CharP.Pi import Mathlib.Algebra.CharP.Quotient +import Mathlib.Algebra.CharP.Reduced import Mathlib.Algebra.CharP.Subring import Mathlib.Algebra.CharP.Two import Mathlib.Algebra.CharZero.Defs @@ -233,6 +234,7 @@ import Mathlib.Algebra.Homology.Functor import Mathlib.Algebra.Homology.HomologicalComplex import Mathlib.Algebra.Homology.HomologicalComplexLimits import Mathlib.Algebra.Homology.Homology +import Mathlib.Algebra.Homology.HomologySequence import Mathlib.Algebra.Homology.Homotopy import Mathlib.Algebra.Homology.HomotopyCategory import Mathlib.Algebra.Homology.HomotopyCategory.HomComplex @@ -805,8 +807,8 @@ import Mathlib.Analysis.NormedSpace.LpEquiv import Mathlib.Analysis.NormedSpace.MStructure import Mathlib.Analysis.NormedSpace.MatrixExponential import Mathlib.Analysis.NormedSpace.MazurUlam -import Mathlib.Analysis.NormedSpace.Multilinear -import Mathlib.Analysis.NormedSpace.MultilinearCurrying +import Mathlib.Analysis.NormedSpace.Multilinear.Basic +import Mathlib.Analysis.NormedSpace.Multilinear.Curry import Mathlib.Analysis.NormedSpace.OperatorNorm import Mathlib.Analysis.NormedSpace.PiLp import Mathlib.Analysis.NormedSpace.Pointwise @@ -1010,6 +1012,7 @@ import Mathlib.CategoryTheory.Generator import Mathlib.CategoryTheory.GlueData import Mathlib.CategoryTheory.GradedObject import Mathlib.CategoryTheory.GradedObject.Bifunctor +import Mathlib.CategoryTheory.GradedObject.Trifunctor import Mathlib.CategoryTheory.Grothendieck import Mathlib.CategoryTheory.Groupoid import Mathlib.CategoryTheory.Groupoid.Basic @@ -2614,6 +2617,7 @@ import Mathlib.MeasureTheory.Measure.Haar.NormedSpace import Mathlib.MeasureTheory.Measure.Haar.OfBasis import Mathlib.MeasureTheory.Measure.Haar.Quotient import Mathlib.MeasureTheory.Measure.Haar.Unique +import Mathlib.MeasureTheory.Measure.HasOuterApproxClosed import Mathlib.MeasureTheory.Measure.Hausdorff import Mathlib.MeasureTheory.Measure.Lebesgue.Basic import Mathlib.MeasureTheory.Measure.Lebesgue.Complex @@ -3119,6 +3123,7 @@ import Mathlib.RingTheory.Valuation.Basic import Mathlib.RingTheory.Valuation.ExtendToLocalization import Mathlib.RingTheory.Valuation.Integers import Mathlib.RingTheory.Valuation.Integral +import Mathlib.RingTheory.Valuation.PrimeMultiplicity import Mathlib.RingTheory.Valuation.Quotient import Mathlib.RingTheory.Valuation.RamificationGroup import Mathlib.RingTheory.Valuation.ValuationRing @@ -3384,7 +3389,7 @@ import Mathlib.Topology.Algebra.InfiniteSum.Order import Mathlib.Topology.Algebra.InfiniteSum.Real import Mathlib.Topology.Algebra.InfiniteSum.Ring import Mathlib.Topology.Algebra.Localization -import Mathlib.Topology.Algebra.Module.Alternating +import Mathlib.Topology.Algebra.Module.Alternating.Basic import Mathlib.Topology.Algebra.Module.Basic import Mathlib.Topology.Algebra.Module.Cardinality import Mathlib.Topology.Algebra.Module.CharacterSpace @@ -3392,7 +3397,7 @@ import Mathlib.Topology.Algebra.Module.Determinant import Mathlib.Topology.Algebra.Module.FiniteDimension import Mathlib.Topology.Algebra.Module.LinearPMap import Mathlib.Topology.Algebra.Module.LocallyConvex -import Mathlib.Topology.Algebra.Module.Multilinear +import Mathlib.Topology.Algebra.Module.Multilinear.Basic import Mathlib.Topology.Algebra.Module.Simple import Mathlib.Topology.Algebra.Module.Star import Mathlib.Topology.Algebra.Module.StrongTopology @@ -3605,6 +3610,7 @@ import Mathlib.Topology.Order.Hom.Esakia import Mathlib.Topology.Order.Lattice import Mathlib.Topology.Order.LowerUpperTopology import Mathlib.Topology.Order.NhdsSet +import Mathlib.Topology.Order.PartialSups import Mathlib.Topology.Order.Priestley import Mathlib.Topology.Order.UpperLowerSetTopology import Mathlib.Topology.Partial diff --git a/Mathlib/Algebra/Algebra/Equiv.lean b/Mathlib/Algebra/Algebra/Equiv.lean index e4ad17ac9fa84..0696d7b8ce404 100644 --- a/Mathlib/Algebra/Algebra/Equiv.lean +++ b/Mathlib/Algebra/Algebra/Equiv.lean @@ -362,7 +362,7 @@ theorem symm_symm (e : A₁ ≃ₐ[R] A₂) : e.symm.symm = e := by #align alg_equiv.symm_symm AlgEquiv.symm_symm theorem symm_bijective : Function.Bijective (symm : (A₁ ≃ₐ[R] A₂) → A₂ ≃ₐ[R] A₁) := - Equiv.bijective ⟨symm, symm, symm_symm, symm_symm⟩ + Function.bijective_iff_has_inverse.mpr ⟨_, symm_symm, symm_symm⟩ #align alg_equiv.symm_bijective AlgEquiv.symm_bijective @[simp] diff --git a/Mathlib/Algebra/Algebra/NonUnitalSubalgebra.lean b/Mathlib/Algebra/Algebra/NonUnitalSubalgebra.lean index 3f902164001e4..9a15a0b11d93f 100644 --- a/Mathlib/Algebra/Algebra/NonUnitalSubalgebra.lean +++ b/Mathlib/Algebra/Algebra/NonUnitalSubalgebra.lean @@ -657,11 +657,11 @@ theorem to_subring_eq_top {R A : Type*} [CommRing R] [Ring A] [Algebra R A] NonUnitalSubalgebra.toNonUnitalSubring_injective.eq_iff' top_toSubring theorem mem_sup_left {S T : NonUnitalSubalgebra R A} : ∀ {x : A}, x ∈ S → x ∈ S ⊔ T := by - rw [←SetLike.le_def] + rw [← SetLike.le_def] exact le_sup_left theorem mem_sup_right {S T : NonUnitalSubalgebra R A} : ∀ {x : A}, x ∈ T → x ∈ S ⊔ T := by - rw [←SetLike.le_def] + rw [← SetLike.le_def] exact le_sup_right theorem mul_mem_sup {S T : NonUnitalSubalgebra R A} {x y : A} (hx : x ∈ S) (hy : y ∈ T) : diff --git a/Mathlib/Algebra/Algebra/Operations.lean b/Mathlib/Algebra/Algebra/Operations.lean index c323c4ccc38f5..05a1580b7614f 100644 --- a/Mathlib/Algebra/Algebra/Operations.lean +++ b/Mathlib/Algebra/Algebra/Operations.lean @@ -636,7 +636,7 @@ instance moduleSet : Module (SetSemiring A) (Submodule R A) where mul_smul s t P := by simp_rw [HSMul.hSMul, SetSemiring.down_mul, ← mul_assoc, span_mul_span] one_smul P := by - simp_rw [HSMul.hSMul, SetSemiring.down_one, ←one_eq_span_one_set, one_mul] + simp_rw [HSMul.hSMul, SetSemiring.down_one, ← one_eq_span_one_set, one_mul] zero_smul P := by simp_rw [HSMul.hSMul, SetSemiring.down_zero, span_empty, bot_mul, bot_eq_zero] smul_zero _ := mul_bot _ diff --git a/Mathlib/Algebra/Algebra/Subalgebra/Unitization.lean b/Mathlib/Algebra/Algebra/Subalgebra/Unitization.lean index 9792a9232f75c..36397a3317628 100644 --- a/Mathlib/Algebra/Algebra/Subalgebra/Unitization.lean +++ b/Mathlib/Algebra/Algebra/Subalgebra/Unitization.lean @@ -131,7 +131,7 @@ theorem _root_.AlgHomClass.unitization_injective' {F R S A : Type*} [CommRing R] (hf : ∀ x : s, f x = x) : Function.Injective f := by refine' (injective_iff_map_eq_zero _).mpr fun x hx => _ induction' x using Unitization.ind with r a - simp_rw [map_add, hf, ←Unitization.algebraMap_eq_inl, AlgHomClass.commutes] at hx + simp_rw [map_add, hf, ← Unitization.algebraMap_eq_inl, AlgHomClass.commutes] at hx rw [add_eq_zero_iff_eq_neg] at hx ⊢ by_cases hr : r = 0 · ext <;> simp [hr] at hx ⊢ diff --git a/Mathlib/Algebra/Associated.lean b/Mathlib/Algebra/Associated.lean index cb3189c57e838..fc1a9c7e8a341 100644 --- a/Mathlib/Algebra/Associated.lean +++ b/Mathlib/Algebra/Associated.lean @@ -779,6 +779,11 @@ theorem quot_mk_eq_mk [Monoid α] (a : α) : Quot.mk Setoid.r a = Associates.mk rfl #align associates.quot_mk_eq_mk Associates.quot_mk_eq_mk +@[simp] +theorem quot_out [Monoid α] (a : Associates α) : Associates.mk (Quot.out a) = a := by + rw [← quot_mk_eq_mk, Quot.out_eq] +#align associates.quot_out Associates.quot_outₓ + theorem forall_associated [Monoid α] {p : Associates α → Prop} : (∀ a, p a) ↔ ∀ a, p (Associates.mk a) := Iff.intro (fun h _ => h _) fun h a => Quotient.inductionOn a h @@ -1158,9 +1163,9 @@ theorem one_or_eq_of_le_of_prime : ∀ p m : Associates α, Prime p → m ≤ p #align associates.one_or_eq_of_le_of_prime Associates.one_or_eq_of_le_of_prime instance : CanonicallyOrderedCommMonoid (Associates α) where - exists_mul_of_le := fun h => h - le_self_mul := fun _ b => ⟨b, rfl⟩ - bot_le := fun _ => one_le + exists_mul_of_le := fun h => h + le_self_mul := fun _ b => ⟨b, rfl⟩ + bot_le := fun _ => one_le theorem dvdNotUnit_iff_lt {a b : Associates α} : DvdNotUnit a b ↔ a < b := dvd_and_not_dvd_iff.symm diff --git a/Mathlib/Algebra/BigOperators/Associated.lean b/Mathlib/Algebra/BigOperators/Associated.lean index 03eb6df430633..6feb4b61756c4 100644 --- a/Mathlib/Algebra/BigOperators/Associated.lean +++ b/Mathlib/Algebra/BigOperators/Associated.lean @@ -135,7 +135,8 @@ theorem finset_prod_mk {p : Finset β} {f : β → α} : -- Porting note: added have : (fun i => Associates.mk (f i)) = Associates.mk ∘ f := funext <| fun x => Function.comp_apply - rw [Finset.prod_eq_multiset_prod, this, ←Multiset.map_map, prod_mk, ←Finset.prod_eq_multiset_prod] + rw [Finset.prod_eq_multiset_prod, this, ← Multiset.map_map, prod_mk, + ← Finset.prod_eq_multiset_prod] #align associates.finset_prod_mk Associates.finset_prod_mk theorem rel_associated_iff_map_eq_map {p q : Multiset α} : diff --git a/Mathlib/Algebra/BigOperators/Basic.lean b/Mathlib/Algebra/BigOperators/Basic.lean index ff5478ad2cda1..9507b0e7a177b 100644 --- a/Mathlib/Algebra/BigOperators/Basic.lean +++ b/Mathlib/Algebra/BigOperators/Basic.lean @@ -216,13 +216,8 @@ theorem map_prod [CommMonoid β] [CommMonoid γ] {G : Type*} [MonoidHomClass G section Deprecated -/-- Deprecated: use `_root_.map_prod` instead. -/ -@[to_additive (attr := deprecated) "Deprecated: use `_root_.map_sum` instead."] -protected theorem MonoidHom.map_prod [CommMonoid β] [CommMonoid γ] (g : β →* γ) (f : α → β) - (s : Finset α) : g (∏ x in s, f x) = ∏ x in s, g (f x) := - map_prod g f s -#align monoid_hom.map_prod MonoidHom.map_prod -#align add_monoid_hom.map_sum AddMonoidHom.map_sum +#align monoid_hom.map_prod map_prodₓ +#align add_monoid_hom.map_sum map_sumₓ /-- Deprecated: use `_root_.map_prod` instead. -/ @[to_additive (attr := deprecated) "Deprecated: use `_root_.map_sum` instead."] @@ -280,7 +275,7 @@ end Deprecated @[to_additive] theorem MonoidHom.coe_finset_prod [MulOneClass β] [CommMonoid γ] (f : α → β →* γ) (s : Finset α) : ⇑(∏ x in s, f x) = ∏ x in s, ⇑(f x) := - (MonoidHom.coeFn β γ).map_prod _ _ + map_prod (MonoidHom.coeFn β γ) _ _ #align monoid_hom.coe_finset_prod MonoidHom.coe_finset_prod #align add_monoid_hom.coe_finset_sum AddMonoidHom.coe_finset_sum @@ -291,7 +286,7 @@ theorem MonoidHom.coe_finset_prod [MulOneClass β] [CommMonoid γ] (f : α → `f : α → β → γ`"] theorem MonoidHom.finset_prod_apply [MulOneClass β] [CommMonoid γ] (f : α → β →* γ) (s : Finset α) (b : β) : (∏ x in s, f x) b = ∏ x in s, f x b := - (MonoidHom.eval b).map_prod _ _ + map_prod (MonoidHom.eval b) _ _ #align monoid_hom.finset_prod_apply MonoidHom.finset_prod_apply #align add_monoid_hom.finset_sum_apply AddMonoidHom.finset_sum_apply @@ -2300,7 +2295,7 @@ end Int @[simp, norm_cast] theorem Units.coe_prod {M : Type*} [CommMonoid M] (f : α → Mˣ) (s : Finset α) : (↑(∏ i in s, f i) : M) = ∏ i in s, (f i : M) := - (Units.coeHom M).map_prod _ _ + map_prod (Units.coeHom M) _ _ #align units.coe_prod Units.coe_prod theorem Units.mk0_prod [CommGroupWithZero β] (s : Finset α) (f : α → β) (h) : diff --git a/Mathlib/Algebra/BigOperators/Finprod.lean b/Mathlib/Algebra/BigOperators/Finprod.lean index 617d15f8172ca..d1f416789aa34 100644 --- a/Mathlib/Algebra/BigOperators/Finprod.lean +++ b/Mathlib/Algebra/BigOperators/Finprod.lean @@ -290,7 +290,7 @@ theorem one_le_finprod' {M : Type*} [OrderedCommMonoid M] {f : α → M} (hf : theorem MonoidHom.map_finprod_plift (f : M →* N) (g : α → M) (h : (mulSupport <| g ∘ PLift.down).Finite) : f (∏ᶠ x, g x) = ∏ᶠ x, f (g x) := by rw [finprod_eq_prod_plift_of_mulSupport_subset h.coe_toFinset.ge, - finprod_eq_prod_plift_of_mulSupport_subset, f.map_prod] + finprod_eq_prod_plift_of_mulSupport_subset, map_prod] rw [h.coe_toFinset] exact mulSupport_comp_subset f.map_one (g ∘ PLift.down) #align monoid_hom.map_finprod_plift MonoidHom.map_finprod_plift @@ -668,7 +668,7 @@ theorem finprod_mem_of_eqOn_one (hf : s.EqOn f 1) : ∏ᶠ i ∈ s, f i = 1 := b "If the product of `f i` over `i ∈ s` is not equal to `0`, then there is some `x ∈ s` such that `f x ≠ 0`."] theorem exists_ne_one_of_finprod_mem_ne_one (h : ∏ᶠ i ∈ s, f i ≠ 1) : ∃ x ∈ s, f x ≠ 1 := by - by_contra' h' + by_contra! h' exact h (finprod_mem_of_eqOn_one h') #align exists_ne_one_of_finprod_mem_ne_one exists_ne_one_of_finprod_mem_ne_one #align exists_ne_zero_of_finsum_mem_ne_zero exists_ne_zero_of_finsum_mem_ne_zero diff --git a/Mathlib/Algebra/BigOperators/Finsupp.lean b/Mathlib/Algebra/BigOperators/Finsupp.lean index c089facb26a1d..315ecf315e71b 100644 --- a/Mathlib/Algebra/BigOperators/Finsupp.lean +++ b/Mathlib/Algebra/BigOperators/Finsupp.lean @@ -314,7 +314,7 @@ namespace Finsupp theorem finset_sum_apply [AddCommMonoid N] (S : Finset ι) (f : ι → α →₀ N) (a : α) : (∑ i in S, f i) a = ∑ i in S, f i a := - (applyAddHom a : (α →₀ N) →+ _).map_sum _ _ + map_sum (applyAddHom a) _ _ #align finsupp.finset_sum_apply Finsupp.finset_sum_apply @[simp] @@ -326,7 +326,7 @@ theorem sum_apply [Zero M] [AddCommMonoid N] {f : α →₀ M} {g : α → M → -- Porting note: inserted ⇑ on the rhs theorem coe_finset_sum [AddCommMonoid N] (S : Finset ι) (f : ι → α →₀ N) : ⇑(∑ i in S, f i) = ∑ i in S, ⇑(f i) := - (coeFnAddHom : (α →₀ N) →+ _).map_sum _ _ + map_sum (coeFnAddHom : (α →₀ N) →+ _) _ _ #align finsupp.coe_finset_sum Finsupp.coe_finset_sum -- Porting note: inserted ⇑ on the rhs @@ -502,7 +502,7 @@ theorem univ_sum_single_apply' [AddCommMonoid M] [Fintype α] (i : α) (m : M) : theorem equivFunOnFinite_symm_eq_sum [Fintype α] [AddCommMonoid M] (f : α → M) : equivFunOnFinite.symm f = ∑ a, Finsupp.single a (f a) := by - rw [←univ_sum_single (equivFunOnFinite.symm f)] + rw [← univ_sum_single (equivFunOnFinite.symm f)] ext simp @@ -587,12 +587,12 @@ theorem support_sum_eq_biUnion {α : Type*} {ι : Type*} {M : Type*} [DecidableE theorem multiset_map_sum [Zero M] {f : α →₀ M} {m : β → γ} {h : α → M → Multiset β} : Multiset.map m (f.sum h) = f.sum fun a b => (h a b).map m := - (Multiset.mapAddMonoidHom m).map_sum _ f.support + map_sum (Multiset.mapAddMonoidHom m) _ f.support #align finsupp.multiset_map_sum Finsupp.multiset_map_sum theorem multiset_sum_sum [Zero M] [AddCommMonoid N] {f : α →₀ M} {h : α → M → Multiset N} : Multiset.sum (f.sum h) = f.sum fun a b => Multiset.sum (h a b) := - (Multiset.sumAddMonoidHom : Multiset N →+ N).map_sum _ f.support + map_sum Multiset.sumAddMonoidHom _ f.support #align finsupp.multiset_sum_sum Finsupp.multiset_sum_sum /-- For disjoint `f1` and `f2`, and function `g`, the product of the products of `g` @@ -669,7 +669,7 @@ lemma sum_cons' [AddCommMonoid M] [AddCommMonoid N] (n : ℕ) (σ : Fin n →₀ end Finsupp theorem Finset.sum_apply' : (∑ k in s, f k) i = ∑ k in s, f k i := - (Finsupp.applyAddHom i : (ι →₀ A) →+ A).map_sum f s + map_sum (Finsupp.applyAddHom i) f s #align finset.sum_apply' Finset.sum_apply' theorem Finsupp.sum_apply' : g.sum k x = g.sum fun i b => k i b x := diff --git a/Mathlib/Algebra/BigOperators/Intervals.lean b/Mathlib/Algebra/BigOperators/Intervals.lean index cc7a28acea0c5..f661f602e41d8 100644 --- a/Mathlib/Algebra/BigOperators/Intervals.lean +++ b/Mathlib/Algebra/BigOperators/Intervals.lean @@ -294,7 +294,7 @@ theorem sum_Ico_by_parts (hmn : m < n) : f (n - 1) • (range n).sum g - f m • (range (m + 1)).sum g + Finset.sum (Ico m (n - 1)) (fun i => f i • (range (i + 1)).sum g - f (i + 1) • (range (i + 1)).sum g) := by - rw [← add_sub, add_comm, ←add_sub, ← sum_sub_distrib] + rw [← add_sub, add_comm, ← add_sub, ← sum_sub_distrib] rw [h₄] have : ∀ i, f i • G (i + 1) - f (i + 1) • G (i + 1) = -((f (i + 1) - f i) • G (i + 1)) := by intro i diff --git a/Mathlib/Algebra/BigOperators/Multiset/Basic.lean b/Mathlib/Algebra/BigOperators/Multiset/Basic.lean index ef6df79f81c8e..ec9c61a8bfc53 100644 --- a/Mathlib/Algebra/BigOperators/Multiset/Basic.lean +++ b/Mathlib/Algebra/BigOperators/Multiset/Basic.lean @@ -310,7 +310,7 @@ theorem prod_map_inv' (m : Multiset α) : (m.map Inv.inv).prod = m.prod⁻¹ := @[to_additive (attr := simp)] theorem prod_map_inv : (m.map fun i => (f i)⁻¹).prod = (m.map f).prod⁻¹ := by -- Porting note: used `convert` - simp_rw [←(m.map f).prod_map_inv', map_map, Function.comp_apply] + simp_rw [← (m.map f).prod_map_inv', map_map, Function.comp_apply] #align multiset.prod_map_inv Multiset.prod_map_inv #align multiset.sum_map_neg Multiset.sum_map_neg diff --git a/Mathlib/Algebra/BigOperators/NatAntidiagonal.lean b/Mathlib/Algebra/BigOperators/NatAntidiagonal.lean index 007de2af7276e..70df8148fbc7e 100644 --- a/Mathlib/Algebra/BigOperators/NatAntidiagonal.lean +++ b/Mathlib/Algebra/BigOperators/NatAntidiagonal.lean @@ -66,9 +66,9 @@ theorem prod_antidiagonal_eq_prod_range_succ_mk {M : Type*} [CommMonoid M] (f : #align finset.nat.sum_antidiagonal_eq_sum_range_succ_mk Finset.Nat.sum_antidiagonal_eq_sum_range_succ_mk /-- This lemma matches more generally than `Finset.Nat.prod_antidiagonal_eq_prod_range_succ_mk` when -using `rw ←`. -/ +using `rw ← `. -/ @[to_additive "This lemma matches more generally than -`Finset.Nat.sum_antidiagonal_eq_sum_range_succ_mk` when using `rw ←`."] +`Finset.Nat.sum_antidiagonal_eq_sum_range_succ_mk` when using `rw ← `."] theorem prod_antidiagonal_eq_prod_range_succ {M : Type*} [CommMonoid M] (f : ℕ → ℕ → M) (n : ℕ) : ∏ ij in antidiagonal n, f ij.1 ij.2 = ∏ k in range n.succ, f k (n - k) := prod_antidiagonal_eq_prod_range_succ_mk _ _ diff --git a/Mathlib/Algebra/BigOperators/Order.lean b/Mathlib/Algebra/BigOperators/Order.lean index cd492c718d79b..2b2a47bc03768 100644 --- a/Mathlib/Algebra/BigOperators/Order.lean +++ b/Mathlib/Algebra/BigOperators/Order.lean @@ -280,7 +280,7 @@ theorem abs_sum_of_nonneg' {G : Type*} [LinearOrderedAddCommGroup G] {f : ι → theorem abs_prod {R : Type*} [LinearOrderedCommRing R] {f : ι → R} {s : Finset ι} : |∏ x in s, f x| = ∏ x in s, |f x| := - (absHom.toMonoidHom : R →* R).map_prod _ _ + map_prod absHom _ _ #align finset.abs_prod Finset.abs_prod section Pigeonhole @@ -545,13 +545,15 @@ variable [DecidableEq ι] @[to_additive] lemma prod_sdiff_le_prod_sdiff : ∏ i in s \ t, f i ≤ ∏ i in t \ s, f i ↔ ∏ i in s, f i ≤ ∏ i in t, f i := by - rw [←mul_le_mul_iff_right, ←prod_union (disjoint_sdiff_inter _ _), sdiff_union_inter, ←prod_union, - inter_comm, sdiff_union_inter]; simpa only [inter_comm] using disjoint_sdiff_inter t s + rw [← mul_le_mul_iff_right, ← prod_union (disjoint_sdiff_inter _ _), sdiff_union_inter, + ← prod_union, inter_comm, sdiff_union_inter]; + simpa only [inter_comm] using disjoint_sdiff_inter t s @[to_additive] lemma prod_sdiff_lt_prod_sdiff : ∏ i in s \ t, f i < ∏ i in t \ s, f i ↔ ∏ i in s, f i < ∏ i in t, f i := by - rw [←mul_lt_mul_iff_right, ←prod_union (disjoint_sdiff_inter _ _), sdiff_union_inter, ←prod_union, - inter_comm, sdiff_union_inter]; simpa only [inter_comm] using disjoint_sdiff_inter t s + rw [← mul_lt_mul_iff_right, ← prod_union (disjoint_sdiff_inter _ _), sdiff_union_inter, + ← prod_union, inter_comm, sdiff_union_inter]; + simpa only [inter_comm] using disjoint_sdiff_inter t s end OrderedCancelCommMonoid @@ -815,10 +817,10 @@ theorem IsAbsoluteValue.abv_sum [Semiring R] [OrderedSemiring S] (abv : R → S) (IsAbsoluteValue.toAbsoluteValue abv).sum_le _ _ #align is_absolute_value.abv_sum IsAbsoluteValue.abv_sum -theorem AbsoluteValue.map_prod [CommSemiring R] [Nontrivial R] [LinearOrderedCommRing S] +nonrec theorem AbsoluteValue.map_prod [CommSemiring R] [Nontrivial R] [LinearOrderedCommRing S] (abv : AbsoluteValue R S) (f : ι → R) (s : Finset ι) : abv (∏ i in s, f i) = ∏ i in s, abv (f i) := - abv.toMonoidHom.map_prod f s + map_prod abv f s #align absolute_value.map_prod AbsoluteValue.map_prod theorem IsAbsoluteValue.map_prod [CommSemiring R] [Nontrivial R] [LinearOrderedCommRing S] diff --git a/Mathlib/Algebra/BigOperators/Pi.lean b/Mathlib/Algebra/BigOperators/Pi.lean index 38adddf53bc5b..8818132d24f12 100644 --- a/Mathlib/Algebra/BigOperators/Pi.lean +++ b/Mathlib/Algebra/BigOperators/Pi.lean @@ -41,7 +41,7 @@ end Pi @[to_additive (attr := simp)] theorem Finset.prod_apply {α : Type*} {β : α → Type*} {γ} [∀ a, CommMonoid (β a)] (a : α) (s : Finset γ) (g : γ → ∀ a, β a) : (∏ c in s, g c) a = ∏ c in s, g c a := - (Pi.evalMonoidHom β a).map_prod _ _ + map_prod (Pi.evalMonoidHom β a) _ _ #align finset.prod_apply Finset.prod_apply #align finset.sum_apply Finset.sum_apply @@ -94,7 +94,7 @@ theorem MonoidHom.functions_ext [Finite I] (G : Type*) [CommMonoid G] (g h : ( (H : ∀ i x, g (Pi.mulSingle i x) = h (Pi.mulSingle i x)) : g = h := by cases nonempty_fintype I ext k - rw [← Finset.univ_prod_mulSingle k, g.map_prod, h.map_prod] + rw [← Finset.univ_prod_mulSingle k, map_prod, map_prod] simp only [H] #align monoid_hom.functions_ext MonoidHom.functions_ext #align add_monoid_hom.functions_ext AddMonoidHom.functions_ext @@ -135,13 +135,13 @@ variable {α β γ : Type*} [CommMonoid α] [CommMonoid β] {s : Finset γ} {f : @[to_additive] theorem fst_prod : (∏ c in s, f c).1 = ∏ c in s, (f c).1 := - (MonoidHom.fst α β).map_prod f s + map_prod (MonoidHom.fst α β) f s #align prod.fst_prod Prod.fst_prod #align prod.fst_sum Prod.fst_sum @[to_additive] theorem snd_prod : (∏ c in s, f c).2 = ∏ c in s, (f c).2 := - (MonoidHom.snd α β).map_prod f s + map_prod (MonoidHom.snd α β) f s #align prod.snd_prod Prod.snd_prod #align prod.snd_sum Prod.snd_sum diff --git a/Mathlib/Algebra/Category/GroupCat/Adjunctions.lean b/Mathlib/Algebra/Category/GroupCat/Adjunctions.lean index bb93f02710d4a..1e2ae1e10e500 100644 --- a/Mathlib/Algebra/Category/GroupCat/Adjunctions.lean +++ b/Mathlib/Algebra/Category/GroupCat/Adjunctions.lean @@ -105,10 +105,10 @@ def free : Type u ⥤ GroupCat where map := FreeGroup.map map_id := by -- This used to be `rw`, but we need `erw` after leanprover/lean4#2644 - intros; ext1; erw [←FreeGroup.map.unique] <;> intros <;> rfl + intros; ext1; erw [← FreeGroup.map.unique] <;> intros <;> rfl map_comp := by -- This used to be `rw`, but we need `erw` after leanprover/lean4#2644 - intros; ext1; erw [←FreeGroup.map.unique] <;> intros <;> rfl + intros; ext1; erw [← FreeGroup.map.unique] <;> intros <;> rfl #align Group.free GroupCat.free /-- The free-forgetful adjunction for groups. diff --git a/Mathlib/Algebra/Category/GroupCat/Colimits.lean b/Mathlib/Algebra/Category/GroupCat/Colimits.lean index e0eb13cbbf456..89441a696195c 100644 --- a/Mathlib/Algebra/Category/GroupCat/Colimits.lean +++ b/Mathlib/Algebra/Category/GroupCat/Colimits.lean @@ -221,7 +221,7 @@ def descMorphism (s : Cocone F) : colimit.{w} F ⟶ s.pt where toFun := descFun F s map_zero' := rfl -- Porting note : in `mathlib3`, nothing needs to be done after `induction` - map_add' x y := Quot.induction_on₂ x y fun _ _ => by dsimp; rw [←quot_add F]; rfl + map_add' x y := Quot.induction_on₂ x y fun _ _ => by dsimp; rw [← quot_add F]; rfl #align AddCommGroup.colimits.desc_morphism AddCommGroupCat.Colimits.descMorphism /-- Evidence that the proposed colimit is the colimit. -/ diff --git a/Mathlib/Algebra/Category/GroupCat/EpiMono.lean b/Mathlib/Algebra/Category/GroupCat/EpiMono.lean index abdd577b5c8f6..8f48effd584a9 100644 --- a/Mathlib/Algebra/Category/GroupCat/EpiMono.lean +++ b/Mathlib/Algebra/Category/GroupCat/EpiMono.lean @@ -318,7 +318,7 @@ theorem comp_eq : (f ≫ show B ⟶ GroupCat.of SX' from g) = f ≫ show B ⟶ G ext a change g (f a) = h (f a) have : f a ∈ { b | h b = g b } := by - rw [←agree] + rw [← agree] use a rw [this] #align Group.surjective_of_epi_auxs.comp_eq GroupCat.SurjectiveOfEpiAuxs.comp_eq diff --git a/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean b/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean index 68367752710b0..a5506c0b063b4 100644 --- a/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean +++ b/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean @@ -455,7 +455,7 @@ def HomEquiv.toRestriction {X Y} (g : Y ⟶ (coextendScalars f).obj X) : (restri rw [← LinearMap.map_smul] erw [smul_eq_mul, mul_one, LinearMap.map_smul] -- Porting note: should probably change CoeFun for obj above - rw [← LinearMap.coe_toAddHom, ←AddHom.toFun_eq_coe] + rw [← LinearMap.coe_toAddHom, ← AddHom.toFun_eq_coe] rw [CoextendScalars.smul_apply (s := f r) (g := g y) (s' := 1), one_mul] simp #align category_theory.Module.restriction_coextension_adj.hom_equiv.to_restriction ModuleCat.RestrictionCoextensionAdj.HomEquiv.toRestriction @@ -496,7 +496,7 @@ protected def unit' : 𝟭 (ModuleCat S) ⟶ restrictScalars f ⋙ coextendScala simp only [ModuleCat.coe_comp, Functor.id_map, Functor.id_obj, Functor.comp_obj, Functor.comp_map] rw [coe_comp, coe_comp, Function.comp, Function.comp] - conv_rhs => rw [← LinearMap.coe_toAddHom, ←AddHom.toFun_eq_coe] + conv_rhs => rw [← LinearMap.coe_toAddHom, ← AddHom.toFun_eq_coe] erw [CoextendScalars.map_apply, AddHom.toFun_eq_coe, LinearMap.coe_toAddHom, restrictScalars.map_apply f] change s • (g y) = g (s • y) @@ -515,9 +515,9 @@ protected def counit' : coextendScalars f ⋙ restrictScalars f ⟶ 𝟭 (Module rw [LinearMap.add_apply] map_smul' := fun r (g : (restrictScalars f).obj ((coextendScalars f).obj X)) => by dsimp - rw [← LinearMap.coe_toAddHom, ←AddHom.toFun_eq_coe] + rw [← LinearMap.coe_toAddHom, ← AddHom.toFun_eq_coe] rw [CoextendScalars.smul_apply (s := f r) (g := g) (s' := 1), one_mul, ← LinearMap.map_smul] - rw [← LinearMap.coe_toAddHom, ←AddHom.toFun_eq_coe] + rw [← LinearMap.coe_toAddHom, ← AddHom.toFun_eq_coe] congr change f r = (f r) • (1 : S) rw [smul_eq_mul (a := f r) (a' := 1), mul_one] } @@ -680,7 +680,7 @@ def Unit.map {X} : X ⟶ (extendScalars f ⋙ restrictScalars f).obj X where map_smul' r x := by letI m1 : Module R S := Module.compHom S f -- Porting note: used to be rfl - dsimp; rw [←TensorProduct.smul_tmul,TensorProduct.smul_tmul'] + dsimp; rw [← TensorProduct.smul_tmul,TensorProduct.smul_tmul'] #align category_theory.Module.extend_restrict_scalars_adj.unit.map ModuleCat.ExtendRestrictScalarsAdj.Unit.map /-- @@ -706,7 +706,7 @@ def Counit.map {Y} : (restrictScalars f ⋙ extendScalars f).obj Y ⟶ Y := by · intros r y dsimp change s • f r • y = f r • s • y - rw [←mul_smul, mul_comm, mul_smul] + rw [← mul_smul, mul_comm, mul_smul] · intros s₁ s₂ ext y change (s₁ + s₂) • y = s₁ • y + s₂ • y diff --git a/Mathlib/Algebra/Category/Ring/Colimits.lean b/Mathlib/Algebra/Category/Ring/Colimits.lean index 89f1a05208d29..74ae83013be3f 100644 --- a/Mathlib/Algebra/Category/Ring/Colimits.lean +++ b/Mathlib/Algebra/Category/Ring/Colimits.lean @@ -310,7 +310,7 @@ def descMorphism (s : Cocone F) : colimit F ⟶ s.pt where map_add' x y := by refine Quot.induction_on₂ x y fun a b => ?_ dsimp [descFun] - rw [←quot_add] + rw [← quot_add] rfl map_mul' x y := by exact Quot.induction_on₂ x y fun a b => rfl #align CommRing.colimits.desc_morphism CommRingCat.Colimits.descMorphism diff --git a/Mathlib/Algebra/Category/Ring/Constructions.lean b/Mathlib/Algebra/Category/Ring/Constructions.lean index ff60abacb7ec2..2e3268cb9e0db 100644 --- a/Mathlib/Algebra/Category/Ring/Constructions.lean +++ b/Mathlib/Algebra/Category/Ring/Constructions.lean @@ -135,7 +135,7 @@ def pushoutCoconeIsColimit : Limits.IsColimit (pushoutCocone f g) := AlgHom.coe_mk, RingHom.coe_mk, MonoidHom.coe_coe, ← eq1, AlgHom.toRingHom_eq_coe, PushoutCocone.ι_app_right, ← eq2, Algebra.TensorProduct.productMap_apply_tmul] change _ = h (a ⊗ₜ 1) * h (1 ⊗ₜ b) - rw [←h.map_mul, Algebra.TensorProduct.tmul_mul_tmul, mul_one, one_mul] + rw [← h.map_mul, Algebra.TensorProduct.tmul_mul_tmul, mul_one, one_mul] rfl set_option linter.uppercaseLean3 false in #align CommRing.pushout_cocone_is_colimit CommRingCat.pushoutCoconeIsColimit @@ -208,7 +208,7 @@ def prodFanIsLimit : IsLimit (prodFan A B) where have eq2 := congr_hom (h ⟨WalkingPair.right⟩) x dsimp at eq1 eq2 -- This used to be `rw`, but we need `erw` after leanprover/lean4#2644 - erw [←eq1, ←eq2] + erw [← eq1, ← eq2] rfl set_option linter.uppercaseLean3 false in #align CommRing.prod_fan_is_limit CommRingCat.prodFanIsLimit @@ -323,7 +323,7 @@ def pullbackConeIsLimit {A B C : CommRingCat.{u}} (f : A ⟶ C) (g : B ⟶ C) : change (m x).1 = (_, _) have eq1 := (congr_arg (fun f : s.pt →+* A => f x) e₁ : _) have eq2 := (congr_arg (fun f : s.pt →+* B => f x) e₂ : _) - rw [←eq1, ←eq2] + rw [← eq1, ← eq2] rfl set_option linter.uppercaseLean3 false in #align CommRing.pullback_cone_is_limit CommRingCat.pullbackConeIsLimit diff --git a/Mathlib/Algebra/CharP/Basic.lean b/Mathlib/Algebra/CharP/Basic.lean index 2dce592d4520c..fd6c8536a0ef8 100644 --- a/Mathlib/Algebra/CharP/Basic.lean +++ b/Mathlib/Algebra/CharP/Basic.lean @@ -5,8 +5,10 @@ Authors: Kenny Lau, Joey van Langen, Casper Putz -/ import Mathlib.Data.Int.ModEq import Mathlib.Data.Nat.Multiplicity +import Mathlib.Data.Nat.Choose.Sum +import Mathlib.Data.Nat.Cast.Prod +import Mathlib.Algebra.Group.ULift import Mathlib.GroupTheory.OrderOfElement -import Mathlib.RingTheory.Nilpotent #align_import algebra.char_p.basic from "leanprover-community/mathlib"@"47a1a73351de8dd6c8d3d32b569c8e434b03ca47" @@ -439,23 +441,6 @@ end CommRing end frobenius -theorem frobenius_inj [CommRing R] [IsReduced R] (p : ℕ) [Fact p.Prime] [CharP R p] : - Function.Injective (frobenius R p) := fun x h H => by - rw [← sub_eq_zero] at H ⊢ - rw [← frobenius_sub] at H - exact IsReduced.eq_zero _ ⟨_, H⟩ -#align frobenius_inj frobenius_inj - -/-- If `ringChar R = 2`, where `R` is a finite reduced commutative ring, -then every `a : R` is a square. -/ -theorem isSquare_of_charTwo' {R : Type*} [Finite R] [CommRing R] [IsReduced R] [CharP R 2] - (a : R) : IsSquare a := by - cases nonempty_fintype R - exact - Exists.imp (fun b h => pow_two b ▸ Eq.symm h) - (((Fintype.bijective_iff_injective_and_card _).mpr ⟨frobenius_inj R 2, rfl⟩).surjective a) -#align is_square_of_char_two' isSquare_of_charTwo' - namespace CharP section @@ -492,23 +477,6 @@ theorem ringChar_zero_iff_CharZero [NonAssocRing R] : ringChar R = 0 ↔ CharZer end -section CommRing - -variable [CommRing R] [IsReduced R] {R} - -@[simp] -theorem pow_prime_pow_mul_eq_one_iff (p k m : ℕ) [Fact p.Prime] [CharP R p] (x : R) : - x ^ (p ^ k * m) = 1 ↔ x ^ m = 1 := by - induction' k with k hk - · rw [pow_zero, one_mul] - · refine' ⟨fun h => _, fun h => _⟩ - · rw [pow_succ, mul_assoc, pow_mul', ← frobenius_def, ← frobenius_one p] at h - exact hk.1 (frobenius_inj R p h) - · rw [pow_mul', h, one_pow] -#align char_p.pow_prime_pow_mul_eq_one_iff CharP.pow_prime_pow_mul_eq_one_iff - -end CommRing - section Semiring variable [NonAssocSemiring R] diff --git a/Mathlib/Algebra/CharP/LocalRing.lean b/Mathlib/Algebra/CharP/LocalRing.lean index a932451a0ca01..c38325cde5363 100644 --- a/Mathlib/Algebra/CharP/LocalRing.lean +++ b/Mathlib/Algebra/CharP/LocalRing.lean @@ -51,8 +51,8 @@ theorem charP_zero_or_prime_power (R : Type*) [CommRing R] [LocalRing R] (q : -- Let `b` be the inverse of `a`. cases' a_unit.exists_left_inv with a_inv h_inv_mul_a have rn_cast_zero : ↑(r ^ n) = (0 : R) := by - rw [← @mul_one R _ ↑(r ^ n), mul_comm, ←Classical.choose_spec a_unit.exists_left_inv, - mul_assoc, ← Nat.cast_mul, ←q_eq_a_mul_rn, CharP.cast_eq_zero R q] + rw [← @mul_one R _ ↑(r ^ n), mul_comm, ← Classical.choose_spec a_unit.exists_left_inv, + mul_assoc, ← Nat.cast_mul, ← q_eq_a_mul_rn, CharP.cast_eq_zero R q] simp have q_eq_rn := Nat.dvd_antisymm ((CharP.cast_eq_zero_iff R q (r ^ n)).mp rn_cast_zero) rn_dvd_q have n_pos : n ≠ 0 := fun n_zero => diff --git a/Mathlib/Algebra/CharP/Quotient.lean b/Mathlib/Algebra/CharP/Quotient.lean index 6de49e382f07b..e72a23cec58e7 100644 --- a/Mathlib/Algebra/CharP/Quotient.lean +++ b/Mathlib/Algebra/CharP/Quotient.lean @@ -5,6 +5,7 @@ Authors: Kenny Lau, Eric Wieser -/ import Mathlib.Algebra.CharP.Basic import Mathlib.RingTheory.Ideal.Quotient +import Mathlib.RingTheory.Ideal.Operations #align_import algebra.char_p.quotient from "leanprover-community/mathlib"@"85e3c05a94b27c84dc6f234cf88326d5e0096ec3" @@ -46,7 +47,7 @@ theorem quotient' {R : Type*} [CommRing R] (p : ℕ) [CharP R p] (I : Ideal R) theorem quotient_iff {R : Type*} [CommRing R] (n : ℕ) [CharP R n] (I : Ideal R) : CharP (R ⧸ I) n ↔ ∀ x : ℕ, ↑x ∈ I → (x : R) = 0 := by refine ⟨fun _ x hx => ?_, CharP.quotient' n I⟩ - rw [CharP.cast_eq_zero_iff R n, ←CharP.cast_eq_zero_iff (R ⧸ I) n _] + rw [CharP.cast_eq_zero_iff R n, ← CharP.cast_eq_zero_iff (R ⧸ I) n _] exact (Submodule.Quotient.mk_eq_zero I).mpr hx /-- `CharP.quotient_iff`, but stated in terms of inclusions of ideals. -/ diff --git a/Mathlib/Algebra/CharP/Reduced.lean b/Mathlib/Algebra/CharP/Reduced.lean new file mode 100644 index 0000000000000..8e7e4eb89e3ed --- /dev/null +++ b/Mathlib/Algebra/CharP/Reduced.lean @@ -0,0 +1,52 @@ +/- +Copyright (c) 2018 Kenny Lau. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kenny Lau, Joey van Langen, Casper Putz +-/ +import Mathlib.Algebra.CharP.Basic +import Mathlib.RingTheory.Nilpotent + +#align_import algebra.char_p.basic from "leanprover-community/mathlib"@"47a1a73351de8dd6c8d3d32b569c8e434b03ca47" + +/-! +# Results about characteristic p reduced rings +-/ + + +open Finset + +open BigOperators + +theorem frobenius_inj (R : Type*) [CommRing R] [IsReduced R] (p : ℕ) [Fact p.Prime] [CharP R p] : + Function.Injective (frobenius R p) := fun x h H => by + rw [← sub_eq_zero] at H ⊢ + rw [← frobenius_sub] at H + exact IsReduced.eq_zero _ ⟨_, H⟩ +#align frobenius_inj frobenius_inj + +/-- If `ringChar R = 2`, where `R` is a finite reduced commutative ring, +then every `a : R` is a square. -/ +theorem isSquare_of_charTwo' {R : Type*} [Finite R] [CommRing R] [IsReduced R] [CharP R 2] + (a : R) : IsSquare a := by + cases nonempty_fintype R + exact + Exists.imp (fun b h => pow_two b ▸ Eq.symm h) + (((Fintype.bijective_iff_injective_and_card _).mpr ⟨frobenius_inj R 2, rfl⟩).surjective a) +#align is_square_of_char_two' isSquare_of_charTwo' + +namespace CharP + +variable {R : Type*} [CommRing R] [IsReduced R] + +@[simp] +theorem pow_prime_pow_mul_eq_one_iff (p k m : ℕ) [Fact p.Prime] [CharP R p] (x : R) : + x ^ (p ^ k * m) = 1 ↔ x ^ m = 1 := by + induction' k with k hk + · rw [pow_zero, one_mul] + · refine' ⟨fun h => _, fun h => _⟩ + · rw [pow_succ, mul_assoc, pow_mul', ← frobenius_def, ← frobenius_one p] at h + exact hk.1 (frobenius_inj R p h) + · rw [pow_mul', h, one_pow] +#align char_p.pow_prime_pow_mul_eq_one_iff CharP.pow_prime_pow_mul_eq_one_iff + +end CharP diff --git a/Mathlib/Algebra/CharZero/Defs.lean b/Mathlib/Algebra/CharZero/Defs.lean index 41a89f75ef662..3653f16065358 100644 --- a/Mathlib/Algebra/CharZero/Defs.lean +++ b/Mathlib/Algebra/CharZero/Defs.lean @@ -132,7 +132,7 @@ instance charZero {M} {n : ℕ} [NeZero n] [AddMonoidWithOne M] [CharZero M] : N instance charZero_one {M} [AddMonoidWithOne M] [CharZero M] : NeZero (1 : M) where out := by - rw [←Nat.cast_one, Nat.cast_ne_zero] + rw [← Nat.cast_one, Nat.cast_ne_zero] trivial instance charZero_ofNat {M} {n : ℕ} [n.AtLeastTwo] [AddMonoidWithOne M] [CharZero M] : diff --git a/Mathlib/Algebra/DirectLimit.lean b/Mathlib/Algebra/DirectLimit.lean index 4016d06a42961..8a95287928f40 100644 --- a/Mathlib/Algebra/DirectLimit.lean +++ b/Mathlib/Algebra/DirectLimit.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2019 Kenny Lau, Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Kenny Lau, Chris Hughes +Authors: Kenny Lau, Chris Hughes, Jujian Zhang -/ import Mathlib.Data.Finset.Order import Mathlib.Algebra.DirectSum.Module @@ -32,7 +32,7 @@ so as to make the operations (addition etc.) "computable". -/ -universe u v w u₁ +universe u v v' v'' w u₁ open Submodule @@ -172,6 +172,94 @@ theorem lift_unique [IsDirected ι (· ≤ ·)] (F : DirectLimit G f →ₗ[R] P · exact DirectLimit.induction_on x fun i x => by rw [lift_of]; rfl #align module.direct_limit.lift_unique Module.DirectLimit.lift_unique +lemma lift_injective [IsDirected ι (· ≤ ·)] + (injective : ∀ i, Function.Injective <| g i) : + Function.Injective (lift R ι G f g Hg) := by + cases isEmpty_or_nonempty ι + · apply Function.injective_of_subsingleton + simp_rw [injective_iff_map_eq_zero] at injective ⊢ + intros z hz + induction' z using DirectLimit.induction_on with _ g + rw [lift_of] at hz + rw [injective _ g hz, _root_.map_zero] + +section functorial + +variable {G' : ι → Type v'} [∀ i, AddCommGroup (G' i)] [∀ i, Module R (G' i)] +variable {f' : ∀ i j, i ≤ j → G' i →ₗ[R] G' j} + +variable {G'' : ι → Type v''} [∀ i, AddCommGroup (G'' i)] [∀ i, Module R (G'' i)] +variable {f'' : ∀ i j, i ≤ j → G'' i →ₗ[R] G'' j} + +/-- +Consider direct limits `lim G` and `lim G'` with direct system `f` and `f'` respectively, any +family of linear maps `gᵢ : Gᵢ ⟶ G'ᵢ` such that `g ∘ f = f' ∘ g` induces a linear map +`lim G ⟶ lim G'`. +-/ +def map (g : (i : ι) → G i →ₗ[R] G' i) (hg : ∀ i j h, g j ∘ₗ f i j h = f' i j h ∘ₗ g i) : + DirectLimit G f →ₗ[R] DirectLimit G' f' := + lift _ _ _ _ (fun i ↦ of _ _ _ _ _ ∘ₗ g i) fun i j h g ↦ by + cases isEmpty_or_nonempty ι + · exact Subsingleton.elim _ _ + · have eq1 := LinearMap.congr_fun (hg i j h) g + simp only [LinearMap.coe_comp, Function.comp_apply] at eq1 ⊢ + rw [eq1, of_f] + +@[simp] lemma map_apply_of (g : (i : ι) → G i →ₗ[R] G' i) + (hg : ∀ i j h, g j ∘ₗ f i j h = f' i j h ∘ₗ g i) + {i : ι} (x : G i) : + map g hg (of _ _ G f _ x) = of R ι G' f' i (g i x) := + lift_of _ _ _ + +@[simp] lemma map_id [IsDirected ι (· ≤ ·)] : + map (fun i ↦ LinearMap.id) (fun _ _ _ ↦ rfl) = LinearMap.id (R := R) (M := DirectLimit G f) := + FunLike.ext _ _ fun x ↦ (isEmpty_or_nonempty ι).elim (fun _ ↦ Subsingleton.elim _ _) fun _ ↦ + x.induction_on fun i g ↦ by simp + +lemma map_comp [IsDirected ι (· ≤ ·)] + (g₁ : (i : ι) → G i →ₗ[R] G' i) (g₂ : (i : ι) → G' i →ₗ[R] G'' i) + (hg₁ : ∀ i j h, g₁ j ∘ₗ f i j h = f' i j h ∘ₗ g₁ i) + (hg₂ : ∀ i j h, g₂ j ∘ₗ f' i j h = f'' i j h ∘ₗ g₂ i) : + (map g₂ hg₂ ∘ₗ map g₁ hg₁ : + DirectLimit G f →ₗ[R] DirectLimit G'' f'') = + (map (fun i ↦ g₂ i ∘ₗ g₁ i) fun i j h ↦ by + rw [LinearMap.comp_assoc, hg₁ i, ← LinearMap.comp_assoc, hg₂ i, LinearMap.comp_assoc] : + DirectLimit G f →ₗ[R] DirectLimit G'' f'') := + FunLike.ext _ _ fun x ↦ (isEmpty_or_nonempty ι).elim (fun _ ↦ Subsingleton.elim _ _) fun _ ↦ + x.induction_on fun i g ↦ by simp + +open LinearEquiv LinearMap in +/-- +Consider direct limits `lim G` and `lim G'` with direct system `f` and `f'` respectively, any +family of equivalences `eᵢ : Gᵢ ≅ G'ᵢ` such that `e ∘ f = f' ∘ e` induces an equivalence +`lim G ≅ lim G'`. +-/ +def congr [IsDirected ι (· ≤ ·)] + (e : (i : ι) → G i ≃ₗ[R] G' i) (he : ∀ i j h, e j ∘ₗ f i j h = f' i j h ∘ₗ e i) : + DirectLimit G f ≃ₗ[R] DirectLimit G' f' := + LinearEquiv.ofLinear (map (e ·) he) + (map (fun i ↦ (e i).symm) fun i j h ↦ by + rw [toLinearMap_symm_comp_eq, ← comp_assoc, he i, comp_assoc, comp_coe, symm_trans_self, + refl_toLinearMap, comp_id]) + (by simp [map_comp]) (by simp [map_comp]) + +lemma congr_apply_of [IsDirected ι (· ≤ ·)] + (e : (i : ι) → G i ≃ₗ[R] G' i) (he : ∀ i j h, e j ∘ₗ f i j h = f' i j h ∘ₗ e i) + {i : ι} (g : G i) : + congr e he (of _ _ G f i g) = of _ _ G' f' i (e i g) := + map_apply_of _ he _ + +open LinearEquiv LinearMap in +lemma congr_symm_apply_of [IsDirected ι (· ≤ ·)] + (e : (i : ι) → G i ≃ₗ[R] G' i) (he : ∀ i j h, e j ∘ₗ f i j h = f' i j h ∘ₗ e i) + {i : ι} (g : G' i) : + (congr e he).symm (of _ _ G' f' i g) = of _ _ G f i ((e i).symm g) := + map_apply_of _ (fun i j h ↦ by + rw [toLinearMap_symm_comp_eq, ← comp_assoc, he i, comp_assoc, comp_coe, symm_trans_self, + refl_toLinearMap, comp_id]) _ + +end functorial + section Totalize open Classical @@ -302,8 +390,8 @@ instance : Inhabited (DirectLimit G f) := instance [IsEmpty ι] : Unique (DirectLimit G f) := Module.DirectLimit.unique _ _ /-- The canonical map from a component to the direct limit. -/ -def of (i) : G i →ₗ[ℤ] DirectLimit G f := - Module.DirectLimit.of ℤ ι G (fun i j hij => (f i j hij).toIntLinearMap) i +def of (i) : G i →+ DirectLimit G f := + (Module.DirectLimit.of ℤ ι G (fun i j hij => (f i j hij).toIntLinearMap) i).toAddMonoidHom #align add_comm_group.direct_limit.of AddCommGroup.DirectLimit.of variable {G f} @@ -337,9 +425,9 @@ variable (G f) /-- The universal property of the direct limit: maps from the components to another abelian group that respect the directed system structure (i.e. make some diagram commute) give rise to a unique map out of the direct limit. -/ -def lift : DirectLimit G f →ₗ[ℤ] P := - Module.DirectLimit.lift ℤ ι G (fun i j hij => (f i j hij).toIntLinearMap) - (fun i => (g i).toIntLinearMap) Hg +def lift : DirectLimit G f →+ P := + (Module.DirectLimit.lift ℤ ι G (fun i j hij => (f i j hij).toIntLinearMap) + (fun i => (g i).toIntLinearMap) Hg).toAddMonoidHom #align add_comm_group.direct_limit.lift AddCommGroup.DirectLimit.lift variable {G f} @@ -350,12 +438,104 @@ theorem lift_of (i x) : lift G f P g Hg (of G f i x) = g i x := #align add_comm_group.direct_limit.lift_of AddCommGroup.DirectLimit.lift_of theorem lift_unique [IsDirected ι (· ≤ ·)] (F : DirectLimit G f →+ P) (x) : - F x = lift G f P (fun i => F.comp (of G f i).toAddMonoidHom) (fun i j hij x => by simp) x := by + F x = lift G f P (fun i => F.comp (of G f i)) (fun i j hij x => by simp) x := by cases isEmpty_or_nonempty ι · simp_rw [Subsingleton.elim x 0, _root_.map_zero] · exact DirectLimit.induction_on x fun i x => by simp #align add_comm_group.direct_limit.lift_unique AddCommGroup.DirectLimit.lift_unique +lemma lift_injective [IsDirected ι (· ≤ ·)] + (injective : ∀ i, Function.Injective <| g i) : + Function.Injective (lift G f P g Hg) := by + cases isEmpty_or_nonempty ι + · apply Function.injective_of_subsingleton + simp_rw [injective_iff_map_eq_zero] at injective ⊢ + intros z hz + induction' z using DirectLimit.induction_on with _ g + rw [lift_of] at hz + rw [injective _ g hz, _root_.map_zero] + +section functorial + +variable {G' : ι → Type v'} [∀ i, AddCommGroup (G' i)] +variable {f' : ∀ i j, i ≤ j → G' i →+ G' j} + +variable {G'' : ι → Type v''} [∀ i, AddCommGroup (G'' i)] +variable {f'' : ∀ i j, i ≤ j → G'' i →+ G'' j} + +/-- +Consider direct limits `lim G` and `lim G'` with direct system `f` and `f'` respectively, any +family of group homomorphisms `gᵢ : Gᵢ ⟶ G'ᵢ` such that `g ∘ f = f' ∘ g` induces a group +homomorphism `lim G ⟶ lim G'`. +-/ +def map (g : (i : ι) → G i →+ G' i) + (hg : ∀ i j h, (g j).comp (f i j h) = (f' i j h).comp (g i)) : + DirectLimit G f →+ DirectLimit G' f' := + lift _ _ _ (fun i ↦ (of _ _ _).comp (g i)) fun i j h g ↦ by + cases isEmpty_or_nonempty ι + · exact Subsingleton.elim _ _ + · have eq1 := FunLike.congr_fun (hg i j h) g + simp only [AddMonoidHom.coe_comp, Function.comp_apply] at eq1 ⊢ + rw [eq1, of_f] + +@[simp] lemma map_apply_of (g : (i : ι) → G i →+ G' i) + (hg : ∀ i j h, (g j).comp (f i j h) = (f' i j h).comp (g i)) + {i : ι} (x : G i) : + map g hg (of G f _ x) = of G' f' i (g i x) := + lift_of _ _ _ _ _ + +@[simp] lemma map_id [IsDirected ι (· ≤ ·)] : + map (fun i ↦ AddMonoidHom.id _) (fun _ _ _ ↦ rfl) = AddMonoidHom.id (DirectLimit G f) := + FunLike.ext _ _ fun x ↦ (isEmpty_or_nonempty ι).elim (fun _ ↦ Subsingleton.elim _ _) fun _ ↦ + x.induction_on fun i g ↦ by simp + +lemma map_comp [IsDirected ι (· ≤ ·)] + (g₁ : (i : ι) → G i →+ G' i) (g₂ : (i : ι) → G' i →+ G'' i) + (hg₁ : ∀ i j h, (g₁ j).comp (f i j h) = (f' i j h).comp (g₁ i)) + (hg₂ : ∀ i j h, (g₂ j).comp (f' i j h) = (f'' i j h).comp (g₂ i)) : + ((map g₂ hg₂).comp (map g₁ hg₁) : + DirectLimit G f →+ DirectLimit G'' f'') = + (map (fun i ↦ (g₂ i).comp (g₁ i)) fun i j h ↦ by + rw [AddMonoidHom.comp_assoc, hg₁ i, ← AddMonoidHom.comp_assoc, hg₂ i, + AddMonoidHom.comp_assoc] : + DirectLimit G f →+ DirectLimit G'' f'') := + FunLike.ext _ _ fun x ↦ (isEmpty_or_nonempty ι).elim (fun _ ↦ Subsingleton.elim _ _) fun _ ↦ + x.induction_on fun i g ↦ by simp + +/-- +Consider direct limits `lim G` and `lim G'` with direct system `f` and `f'` respectively, any +family of equivalences `eᵢ : Gᵢ ≅ G'ᵢ` such that `e ∘ f = f' ∘ e` induces an equivalence +`lim G ⟶ lim G'`. +-/ +def congr [IsDirected ι (· ≤ ·)] + (e : (i : ι) → G i ≃+ G' i) + (he : ∀ i j h, (e j).toAddMonoidHom.comp (f i j h) = (f' i j h).comp (e i)) : + DirectLimit G f ≃+ DirectLimit G' f' := + AddMonoidHom.toAddEquiv (map (e ·) he) + (map (fun i ↦ (e i).symm) fun i j h ↦ FunLike.ext _ _ fun x ↦ by + have eq1 := FunLike.congr_fun (he i j h) ((e i).symm x) + simp only [AddMonoidHom.coe_comp, AddEquiv.coe_toAddMonoidHom, Function.comp_apply, + AddMonoidHom.coe_coe, AddEquiv.apply_symm_apply] at eq1 ⊢ + simp [← eq1, of_f]) + (by rw [map_comp]; convert map_id <;> aesop) + (by rw [map_comp]; convert map_id <;> aesop) + +lemma congr_apply_of [IsDirected ι (· ≤ ·)] + (e : (i : ι) → G i ≃+ G' i) + (he : ∀ i j h, (e j).toAddMonoidHom.comp (f i j h) = (f' i j h).comp (e i)) + {i : ι} (g : G i) : + congr e he (of G f i g) = of G' f' i (e i g) := + map_apply_of _ he _ + +lemma congr_symm_apply_of [IsDirected ι (· ≤ ·)] + (e : (i : ι) → G i ≃+ G' i) + (he : ∀ i j h, (e j).toAddMonoidHom.comp (f i j h) = (f' i j h).comp (e i)) + {i : ι} (g : G' i) : + (congr e he).symm (of G' f' i g) = of G f i ((e i).symm g) := by + simp only [congr, AddMonoidHom.toAddEquiv_symm_apply, map_apply_of, AddMonoidHom.coe_coe] + +end functorial + end DirectLimit end AddCommGroup @@ -690,6 +870,95 @@ theorem lift_unique [IsDirected ι (· ≤ ·)] (F : DirectLimit G f →+* P) (x · exact DirectLimit.induction_on x fun i x => by simp [lift_of] #align ring.direct_limit.lift_unique Ring.DirectLimit.lift_unique +lemma lift_injective [Nonempty ι] [IsDirected ι (· ≤ ·)] + (injective : ∀ i, Function.Injective <| g i) : + Function.Injective (lift G f P g Hg) := by + simp_rw [injective_iff_map_eq_zero] at injective ⊢ + intros z hz + induction' z using DirectLimit.induction_on with _ g + rw [lift_of] at hz + rw [injective _ g hz, _root_.map_zero] + +section functorial + +variable {f : ∀ i j, i ≤ j → G i →+* G j} +variable {G' : ι → Type v'} [∀ i, CommRing (G' i)] +variable {f' : ∀ i j, i ≤ j → G' i →+* G' j} + +variable {G'' : ι → Type v''} [∀ i, CommRing (G'' i)] +variable {f'' : ∀ i j, i ≤ j → G'' i →+* G'' j} + +variable [Nonempty ι] +/-- +Consider direct limits `lim G` and `lim G'` with direct system `f` and `f'` respectively, any +family of ring homomorphisms `gᵢ : Gᵢ ⟶ G'ᵢ` such that `g ∘ f = f' ∘ g` induces a ring +homomorphism `lim G ⟶ lim G'`. +-/ +def map (g : (i : ι) → G i →+* G' i) + (hg : ∀ i j h, (g j).comp (f i j h) = (f' i j h).comp (g i)) : + DirectLimit G (fun _ _ h ↦ f _ _ h) →+* DirectLimit G' fun _ _ h ↦ f' _ _ h := + lift _ _ _ (fun i ↦ (of _ _ _).comp (g i)) fun i j h g ↦ by + have eq1 := FunLike.congr_fun (hg i j h) g + simp only [RingHom.coe_comp, Function.comp_apply] at eq1 ⊢ + rw [eq1, of_f] + +@[simp] lemma map_apply_of (g : (i : ι) → G i →+* G' i) + (hg : ∀ i j h, (g j).comp (f i j h) = (f' i j h).comp (g i)) + {i : ι} (x : G i) : + map g hg (of G _ _ x) = of G' (fun _ _ h ↦ f' _ _ h) i (g i x) := + lift_of _ _ _ _ _ + +@[simp] lemma map_id [IsDirected ι (· ≤ ·)] : + map (fun i ↦ RingHom.id _) (fun _ _ _ ↦ rfl) = + RingHom.id (DirectLimit G fun _ _ h ↦ f _ _ h) := + FunLike.ext _ _ fun x ↦ x.induction_on fun i g ↦ by simp + +lemma map_comp [IsDirected ι (· ≤ ·)] + (g₁ : (i : ι) → G i →+* G' i) (g₂ : (i : ι) → G' i →+* G'' i) + (hg₁ : ∀ i j h, (g₁ j).comp (f i j h) = (f' i j h).comp (g₁ i)) + (hg₂ : ∀ i j h, (g₂ j).comp (f' i j h) = (f'' i j h).comp (g₂ i)) : + ((map g₂ hg₂).comp (map g₁ hg₁) : + DirectLimit G (fun _ _ h ↦ f _ _ h) →+* DirectLimit G'' fun _ _ h ↦ f'' _ _ h) = + (map (fun i ↦ (g₂ i).comp (g₁ i)) fun i j h ↦ by + rw [RingHom.comp_assoc, hg₁ i, ← RingHom.comp_assoc, hg₂ i, RingHom.comp_assoc] : + DirectLimit G (fun _ _ h ↦ f _ _ h) →+* DirectLimit G'' fun _ _ h ↦ f'' _ _ h) := + FunLike.ext _ _ fun x ↦ x.induction_on fun i g ↦ by simp + +/-- +Consider direct limits `lim G` and `lim G'` with direct system `f` and `f'` respectively, any +family of equivalences `eᵢ : Gᵢ ≅ G'ᵢ` such that `e ∘ f = f' ∘ e` induces an equivalence +`lim G ⟶ lim G'`. +-/ +def congr [IsDirected ι (· ≤ ·)] + (e : (i : ι) → G i ≃+* G' i) + (he : ∀ i j h, (e j).toRingHom.comp (f i j h) = (f' i j h).comp (e i)) : + DirectLimit G (fun _ _ h ↦ f _ _ h) ≃+* DirectLimit G' fun _ _ h ↦ f' _ _ h := + RingEquiv.ofHomInv + (map (e ·) he) + (map (fun i ↦ (e i).symm) fun i j h ↦ FunLike.ext _ _ fun x ↦ by + have eq1 := FunLike.congr_fun (he i j h) ((e i).symm x) + simp only [RingEquiv.toRingHom_eq_coe, RingHom.coe_comp, RingHom.coe_coe, Function.comp_apply, + RingEquiv.apply_symm_apply] at eq1 ⊢ + simp [← eq1, of_f]) + (FunLike.ext _ _ fun x ↦ x.induction_on <| by simp) + (FunLike.ext _ _ fun x ↦ x.induction_on <| by simp) + +lemma congr_apply_of [IsDirected ι (· ≤ ·)] + (e : (i : ι) → G i ≃+* G' i) + (he : ∀ i j h, (e j).toRingHom.comp (f i j h) = (f' i j h).comp (e i)) + {i : ι} (g : G i) : + congr e he (of G _ i g) = of G' (fun _ _ h ↦ f' _ _ h) i (e i g) := + map_apply_of _ he _ + +lemma congr_symm_apply_of [IsDirected ι (· ≤ ·)] + (e : (i : ι) → G i ≃+* G' i) + (he : ∀ i j h, (e j).toRingHom.comp (f i j h) = (f' i j h).comp (e i)) + {i : ι} (g : G' i) : + (congr e he).symm (of G' _ i g) = of G (fun _ _ h ↦ f _ _ h) i ((e i).symm g) := by + simp only [congr, RingEquiv.ofHomInv_symm_apply, map_apply_of, RingHom.coe_coe] + +end functorial + end DirectLimit end diff --git a/Mathlib/Algebra/DirectSum/Algebra.lean b/Mathlib/Algebra/DirectSum/Algebra.lean index 034933cc16d76..e1e9c2b14c5b0 100644 --- a/Mathlib/Algebra/DirectSum/Algebra.lean +++ b/Mathlib/Algebra/DirectSum/Algebra.lean @@ -62,13 +62,13 @@ instance _root_.GradedMonoid.smulCommClass_right : SMulCommClass R (GradedMonoid A) (GradedMonoid A) where smul_comm s x y := by dsimp - rw [GAlgebra.smul_def, GAlgebra.smul_def, ←mul_assoc, GAlgebra.commutes, mul_assoc] + rw [GAlgebra.smul_def, GAlgebra.smul_def, ← mul_assoc, GAlgebra.commutes, mul_assoc] instance _root_.GradedMonoid.isScalarTower_right : IsScalarTower R (GradedMonoid A) (GradedMonoid A) where smul_assoc s x y := by dsimp - rw [GAlgebra.smul_def, GAlgebra.smul_def, ←mul_assoc, GAlgebra.commutes, mul_assoc] + rw [GAlgebra.smul_def, GAlgebra.smul_def, ← mul_assoc, GAlgebra.commutes, mul_assoc] instance : Algebra R (⨁ i, A i) where toFun := (DirectSum.of A 0).comp GAlgebra.toFun @@ -121,7 +121,7 @@ def toAlgebra (f : ∀ i, A i →ₗ[R] B) (hone : f _ GradedMonoid.GOne.one = 1 commutes' := fun r => by show toModule R _ _ f (algebraMap R _ r) = _ rw [Algebra.algebraMap_eq_smul_one, Algebra.algebraMap_eq_smul_one, map_smul, one_def, - ←lof_eq_of R, toModule_lof, hone] } + ← lof_eq_of R, toModule_lof, hone] } #align direct_sum.to_algebra DirectSum.toAlgebra /-- Two `AlgHom`s out of a direct sum are equal if they agree on the generators. diff --git a/Mathlib/Algebra/DualNumber.lean b/Mathlib/Algebra/DualNumber.lean index c8b6f1682298f..03a90e2a4815b 100644 --- a/Mathlib/Algebra/DualNumber.lean +++ b/Mathlib/Algebra/DualNumber.lean @@ -141,7 +141,7 @@ def lift : (fg.val.1, fg.val.2 1), fg.prop.1 _ _, fun a => show fg.val.2 1 * fg.val.1 a = fg.val.1 a * fg.val.2 1 by - rw [←fg.prop.2.1, ←fg.prop.2.2, smul_eq_mul, op_smul_eq_mul, mul_one, one_mul]⟩ + rw [← fg.prop.2.1, ← fg.prop.2.2, smul_eq_mul, op_smul_eq_mul, mul_one, one_mul]⟩ left_inv := fun fe => Subtype.ext <| Prod.ext rfl <| show fe.val.1 1 * fe.val.2 = fe.val.2 by rw [map_one, one_mul] @@ -164,12 +164,12 @@ theorem lift_apply_apply (fe : {_fe : (A →ₐ[R] B) × B // _}) (a : A[ε]) : /-- Scaling on the left is sent by `DualNumber.lift` to multiplication on the left -/ @[simp] theorem lift_smul (fe : {fe : (A →ₐ[R] B) × B // _}) (a : A) (ad : A[ε]) : lift fe (a • ad) = fe.val.1 a * lift fe ad := by - rw [←inl_mul_eq_smul, map_mul, lift_apply_inl] + rw [← inl_mul_eq_smul, map_mul, lift_apply_inl] /-- Scaling on the right is sent by `DualNumber.lift` to multiplication on the right -/ @[simp] theorem lift_op_smul (fe : {fe : (A →ₐ[R] B) × B // _}) (a : A) (ad : A[ε]) : lift fe (MulOpposite.op a • ad) = lift fe ad * fe.val.1 a := by - rw [←mul_inl_eq_op_smul, map_mul, lift_apply_inl] + rw [← mul_inl_eq_op_smul, map_mul, lift_apply_inl] /-- When applied to `ε`, `DualNumber.lift` produces the element of `B` that squares to 0. -/ @[simp] theorem lift_apply_eps diff --git a/Mathlib/Algebra/EuclideanDomain/Instances.lean b/Mathlib/Algebra/EuclideanDomain/Instances.lean index 6a40763a50706..87c43b6a42342 100644 --- a/Mathlib/Algebra/EuclideanDomain/Instances.lean +++ b/Mathlib/Algebra/EuclideanDomain/Instances.lean @@ -32,7 +32,7 @@ instance Int.euclideanDomain : EuclideanDomain ℤ := mul_left_not_lt := fun a b b0 => not_lt_of_ge <| by rw [← mul_one a.natAbs, Int.natAbs_mul] - rw [←Int.natAbs_pos] at b0 + rw [← Int.natAbs_pos] at b0 exact Nat.mul_le_mul_of_nonneg_left b0 } -- see Note [lower instance priority] diff --git a/Mathlib/Algebra/FreeAlgebra.lean b/Mathlib/Algebra/FreeAlgebra.lean index c76ee49f985fa..5e40bc2e7320a 100644 --- a/Mathlib/Algebra/FreeAlgebra.lean +++ b/Mathlib/Algebra/FreeAlgebra.lean @@ -239,7 +239,7 @@ instance instAddCommMonoid : AddCommMonoid (FreeAlgebra R X) where nsmul_succ n := by rintro ⟨a⟩ dsimp only [HSMul.hSMul, instSMul, Quot.map] - rw [map_add, map_one, add_comm, mk_mul, mk_mul, ←one_add_mul (_ : FreeAlgebra R X)] + rw [map_add, map_one, add_comm, mk_mul, mk_mul, ← one_add_mul (_ : FreeAlgebra R X)] congr 1 exact Quot.sound Rel.add_scalar @@ -276,10 +276,10 @@ instance {R S A} [CommSemiring R] [CommSemiring S] [CommSemiring A] IsScalarTower R S (FreeAlgebra A X) where smul_assoc r s x := by change algebraMap S A (r • s) • x = algebraMap R A _ • (algebraMap S A _ • x) - rw [←smul_assoc] + rw [← smul_assoc] congr simp only [Algebra.algebraMap_eq_smul_one, smul_eq_mul] - rw [smul_assoc, ←smul_one_mul] + rw [smul_assoc, ← smul_one_mul] instance {R S A} [CommSemiring R] [CommSemiring S] [CommSemiring A] [Algebra R A] [Algebra S A] : SMulCommClass R S (FreeAlgebra A X) where @@ -593,7 +593,7 @@ variable {A : Type*} [Semiring A] [Algebra R A] /-- Noncommutative version of `Algebra.adjoin_range_eq_range_aeval`. -/ theorem _root_.Algebra.adjoin_range_eq_range_freeAlgebra_lift (f : X → A) : Algebra.adjoin R (Set.range f) = (FreeAlgebra.lift R f).range := by - simp only [← Algebra.map_top, ←adjoin_range_ι, AlgHom.map_adjoin, ← Set.range_comp, + simp only [← Algebra.map_top, ← adjoin_range_ι, AlgHom.map_adjoin, ← Set.range_comp, (· ∘ ·), lift_ι_apply] /-- Noncommutative version of `Algebra.adjoin_range_eq_range`. -/ diff --git a/Mathlib/Algebra/Group/Basic.lean b/Mathlib/Algebra/Group/Basic.lean index c2af1e17fc9e1..6aca743a6ae77 100644 --- a/Mathlib/Algebra/Group/Basic.lean +++ b/Mathlib/Algebra/Group/Basic.lean @@ -893,7 +893,7 @@ theorem div_eq_of_eq_mul' {a b c : G} (h : a = b * c) : a / b = c := by @[to_additive (attr := simp)] theorem mul_div_mul_left_eq_div (a b c : G) : c * a / (c * b) = a / b := by - rw [div_eq_mul_inv, mul_inv_rev, mul_comm b⁻¹ c⁻¹, mul_comm c a, mul_assoc, ←mul_assoc c, + rw [div_eq_mul_inv, mul_inv_rev, mul_comm b⁻¹ c⁻¹, mul_comm c a, mul_assoc, ← mul_assoc c, mul_right_inv, one_mul, div_eq_mul_inv] #align mul_div_mul_left_eq_div mul_div_mul_left_eq_div #align add_sub_add_left_eq_sub add_sub_add_left_eq_sub diff --git a/Mathlib/Algebra/Group/Commute/Defs.lean b/Mathlib/Algebra/Group/Commute/Defs.lean index a8d50465e2c69..0cf1cc337e983 100644 --- a/Mathlib/Algebra/Group/Commute/Defs.lean +++ b/Mathlib/Algebra/Group/Commute/Defs.lean @@ -30,7 +30,7 @@ Most of the proofs come from the properties of `SemiconjBy`. variable {G : Type*} /-- Two elements commute if `a * b = b * a`. -/ -@[to_additive AddCommute "Two elements additively commute if `a + b = b + a`"] +@[to_additive "Two elements additively commute if `a + b = b + a`"] def Commute {S : Type*} [Mul S] (a b : S) : Prop := SemiconjBy a b b #align commute Commute @@ -67,7 +67,7 @@ protected theorem symm {a b : S} (h : Commute a b) : Commute b a := protected theorem semiconjBy {a b : S} (h : Commute a b) : SemiconjBy a b b := h #align commute.semiconj_by Commute.semiconjBy -#align add_commute.semiconj_by AddCommute.semiconjBy +#align add_commute.semiconj_by AddCommute.addSemiconjBy @[to_additive] protected theorem symm_iff {a b : S} : Commute a b ↔ Commute b a := @@ -133,7 +133,7 @@ protected theorem mul_mul_mul_comm (hbc : Commute b c) (a d : S) : end Semigroup @[to_additive] -protected theorem all {S : Type*} [CommSemigroup S] (a b : S) : Commute a b := +protected theorem all {S : Type*} [CommMagma S] (a b : S) : Commute a b := mul_comm a b #align commute.all Commute.allₓ #align add_commute.all AddCommute.allₓ diff --git a/Mathlib/Algebra/Group/Defs.lean b/Mathlib/Algebra/Group/Defs.lean index 0933d0131d034..dc0da7bc252e4 100644 --- a/Mathlib/Algebra/Group/Defs.lean +++ b/Mathlib/Algebra/Group/Defs.lean @@ -315,87 +315,86 @@ instance Semigroup.to_isAssociative : IsAssociative G (· * ·) := end Semigroup -/-- A commutative semigroup is a type with an associative commutative `(*)`. -/ +/-- A commutative addition is a type with an addition which commutes-/ +@[ext] +class AddCommMagma (G : Type u) extends Add G where + /-- Addition is commutative in an additive commutative semigroup. -/ + protected add_comm : ∀ a b : G, a + b = b + a + +/-- A commutative multiplication is a type with a multiplication which commutes-/ @[ext] -class CommSemigroup (G : Type u) extends Semigroup G where +class CommMagma (G : Type u) extends Mul G where /-- Multiplication is commutative in a commutative semigroup. -/ protected mul_comm : ∀ a b : G, a * b = b * a + +attribute [to_additive] CommMagma + +/-- A commutative semigroup is a type with an associative commutative `(*)`. -/ +@[ext] +class CommSemigroup (G : Type u) extends Semigroup G, CommMagma G where #align comm_semigroup CommSemigroup #align comm_semigroup.ext_iff CommSemigroup.ext_iff #align comm_semigroup.ext CommSemigroup.ext /-- A commutative additive semigroup is a type with an associative commutative `(+)`. -/ @[ext] -class AddCommSemigroup (G : Type u) extends AddSemigroup G where - /-- Addition is commutative in an additive commutative semigroup. -/ - protected add_comm : ∀ a b : G, a + b = b + a +class AddCommSemigroup (G : Type u) extends AddSemigroup G, AddCommMagma G where #align add_comm_semigroup AddCommSemigroup #align add_comm_semigroup.ext AddCommSemigroup.ext #align add_comm_semigroup.ext_iff AddCommSemigroup.ext_iff attribute [to_additive] CommSemigroup +attribute [to_additive existing] CommSemigroup.toCommMagma -section CommSemigroup +section CommMagma -variable [CommSemigroup G] +variable [CommMagma G] @[to_additive] -theorem mul_comm : ∀ a b : G, a * b = b * a := - CommSemigroup.mul_comm +theorem mul_comm : ∀ a b : G, a * b = b * a := CommMagma.mul_comm #align mul_comm mul_comm #align add_comm add_comm @[to_additive] -instance CommSemigroup.to_isCommutative : IsCommutative G (· * ·) := - ⟨mul_comm⟩ -#align comm_semigroup.to_is_commutative CommSemigroup.to_isCommutative -#align add_comm_semigroup.to_is_commutative AddCommSemigroup.to_isCommutative - -/-- Any `CommSemigroup G` that satisfies `IsRightCancelMul G` also satisfies -`IsLeftCancelMul G`. -/ -@[to_additive AddCommSemigroup.IsRightCancelAdd.toIsLeftCancelAdd "Any -`AddCommSemigroup G` that satisfies `IsRightCancelAdd G` also satisfies -`IsLeftCancelAdd G`."] -lemma CommSemigroup.IsRightCancelMul.toIsLeftCancelMul (G : Type u) [CommSemigroup G] - [IsRightCancelMul G] : IsLeftCancelMul G := +instance CommMagma.to_isCommutative : IsCommutative G (· * ·) := ⟨mul_comm⟩ +#align comm_semigroup.to_is_commutative CommMagma.to_isCommutative +#align add_comm_semigroup.to_is_commutative AddCommMagma.to_isCommutative + +/-- Any `CommMagma G` that satisfies `IsRightCancelMul G` also satisfies `IsLeftCancelMul G`. -/ +@[to_additive AddCommMagma.IsRightCancelAdd.toIsLeftCancelAdd "Any `AddCommMagma G` that satisfies +`IsRightCancelAdd G` also satisfies `IsLeftCancelAdd G`."] +lemma CommMagma.IsRightCancelMul.toIsLeftCancelMul (G : Type u) [CommMagma G] [IsRightCancelMul G] : + IsLeftCancelMul G := ⟨fun _ _ _ h => mul_right_cancel <| (mul_comm _ _).trans (h.trans (mul_comm _ _))⟩ -#align comm_semigroup.is_right_cancel_mul.to_is_left_cancel_mul CommSemigroup.IsRightCancelMul.toIsLeftCancelMul -#align add_comm_semigroup.is_right_cancel_add.to_is_left_cancel_add AddCommSemigroup.IsRightCancelAdd.toIsLeftCancelAdd - -/-- Any `CommSemigroup G` that satisfies `IsLeftCancelMul G` also satisfies -`IsRightCancelMul G`. -/ -@[to_additive AddCommSemigroup.IsLeftCancelAdd.toIsRightCancelAdd "Any -`AddCommSemigroup G` that satisfies `IsLeftCancelAdd G` also satisfies -`IsRightCancelAdd G`."] -lemma CommSemigroup.IsLeftCancelMul.toIsRightCancelMul (G : Type u) [CommSemigroup G] - [IsLeftCancelMul G] : IsRightCancelMul G := +#align comm_semigroup.is_right_cancel_mul.to_is_left_cancel_mul CommMagma.IsRightCancelMul.toIsLeftCancelMul +#align add_comm_semigroup.is_right_cancel_add.to_is_left_cancel_add AddCommMagma.IsRightCancelAdd.toIsLeftCancelAdd + +/-- Any `CommMagma G` that satisfies `IsLeftCancelMul G` also satisfies `IsRightCancelMul G`. -/ +@[to_additive AddCommMagma.IsLeftCancelAdd.toIsRightCancelAdd "Any `AddCommMagma G` that satisfies +`IsLeftCancelAdd G` also satisfies `IsRightCancelAdd G`."] +lemma CommMagma.IsLeftCancelMul.toIsRightCancelMul (G : Type u) [CommMagma G] [IsLeftCancelMul G] : + IsRightCancelMul G := ⟨fun _ _ _ h => mul_left_cancel <| (mul_comm _ _).trans (h.trans (mul_comm _ _))⟩ -#align comm_semigroup.is_left_cancel_mul.to_is_right_cancel_mul CommSemigroup.IsLeftCancelMul.toIsRightCancelMul -#align add_comm_semigroup.is_left_cancel_add.to_is_right_cancel_add AddCommSemigroup.IsLeftCancelAdd.toIsRightCancelAdd - -/-- Any `CommSemigroup G` that satisfies `IsLeftCancelMul G` also satisfies -`IsCancelMul G`. -/ -@[to_additive AddCommSemigroup.IsLeftCancelAdd.toIsCancelAdd "Any -`AddCommSemigroup G` that satisfies `IsLeftCancelAdd G` also satisfies -`IsCancelAdd G`."] -lemma CommSemigroup.IsLeftCancelMul.toIsCancelMul (G : Type u) [CommSemigroup G] - [IsLeftCancelMul G] : IsCancelMul G := - { CommSemigroup.IsLeftCancelMul.toIsRightCancelMul G with } -#align comm_semigroup.is_left_cancel_mul.to_is_cancel_mul CommSemigroup.IsLeftCancelMul.toIsCancelMul -#align add_comm_semigroup.is_left_cancel_add.to_is_cancel_add AddCommSemigroup.IsLeftCancelAdd.toIsCancelAdd - -/-- Any `CommSemigroup G` that satisfies `IsRightCancelMul G` also satisfies -`IsCancelMul G`. -/ -@[to_additive AddCommSemigroup.IsRightCancelAdd.toIsCancelAdd "Any -`AddCommSemigroup G` that satisfies `IsRightCancelAdd G` also satisfies -`IsCancelAdd G`."] -lemma CommSemigroup.IsRightCancelMul.toIsCancelMul (G : Type u) [CommSemigroup G] - [IsRightCancelMul G] : IsCancelMul G := - { CommSemigroup.IsRightCancelMul.toIsLeftCancelMul G with } -#align comm_semigroup.is_right_cancel_mul.to_is_cancel_mul CommSemigroup.IsRightCancelMul.toIsCancelMul -#align add_comm_semigroup.is_right_cancel_add.to_is_cancel_add AddCommSemigroup.IsRightCancelAdd.toIsCancelAdd - -end CommSemigroup +#align comm_semigroup.is_left_cancel_mul.to_is_right_cancel_mul CommMagma.IsLeftCancelMul.toIsRightCancelMul +#align add_comm_semigroup.is_left_cancel_add.to_is_right_cancel_add AddCommMagma.IsLeftCancelAdd.toIsRightCancelAdd + +/-- Any `CommMagma G` that satisfies `IsLeftCancelMul G` also satisfies `IsCancelMul G`. -/ +@[to_additive AddCommMagma.IsLeftCancelAdd.toIsCancelAdd "Any `AddCommMagma G` that satisfies +`IsLeftCancelAdd G` also satisfies `IsCancelAdd G`."] +lemma CommMagma.IsLeftCancelMul.toIsCancelMul (G : Type u) [CommMagma G] [IsLeftCancelMul G] : + IsCancelMul G := { CommMagma.IsLeftCancelMul.toIsRightCancelMul G with } +#align comm_semigroup.is_left_cancel_mul.to_is_cancel_mul CommMagma.IsLeftCancelMul.toIsCancelMul +#align add_comm_semigroup.is_left_cancel_add.to_is_cancel_add AddCommMagma.IsLeftCancelAdd.toIsCancelAdd + +/-- Any `CommMagma G` that satisfies `IsRightCancelMul G` also satisfies `IsCancelMul G`. -/ +@[to_additive AddCommMagma.IsRightCancelAdd.toIsCancelAdd "Any `AddCommMagma G` that satisfies +`IsRightCancelAdd G` also satisfies `IsCancelAdd G`."] +lemma CommMagma.IsRightCancelMul.toIsCancelMul (G : Type u) [CommMagma G] [IsRightCancelMul G] : + IsCancelMul G := { CommMagma.IsRightCancelMul.toIsLeftCancelMul G with } +#align comm_semigroup.is_right_cancel_mul.to_is_cancel_mul CommMagma.IsRightCancelMul.toIsCancelMul +#align add_comm_semigroup.is_right_cancel_add.to_is_cancel_add AddCommMagma.IsRightCancelAdd.toIsCancelAdd + +end CommMagma /-- A `LeftCancelSemigroup` is a semigroup such that `a * b = a * c` implies `b = c`. -/ @[ext] @@ -789,7 +788,7 @@ attribute [to_additive existing] CancelCommMonoid.toCommMonoid @[to_additive] instance (priority := 100) CancelCommMonoid.toCancelMonoid (M : Type u) [CancelCommMonoid M] : CancelMonoid M := - { CommSemigroup.IsLeftCancelMul.toIsRightCancelMul M with } + { CommMagma.IsLeftCancelMul.toIsRightCancelMul M with } #align cancel_comm_monoid.to_cancel_monoid CancelCommMonoid.toCancelMonoid #align add_cancel_comm_monoid.to_cancel_add_monoid AddCancelCommMonoid.toAddCancelMonoid diff --git a/Mathlib/Algebra/Group/Equiv/Basic.lean b/Mathlib/Algebra/Group/Equiv/Basic.lean index 3964a3922a56f..91c4988eb003e 100644 --- a/Mathlib/Algebra/Group/Equiv/Basic.lean +++ b/Mathlib/Algebra/Group/Equiv/Basic.lean @@ -343,8 +343,8 @@ theorem symm_symm (f : M ≃* N) : f.symm.symm = f := rfl #align add_equiv.symm_symm AddEquiv.symm_symm @[to_additive] -theorem symm_bijective : Function.Bijective (symm : M ≃* N → N ≃* M) := - Equiv.bijective ⟨symm, symm, symm_symm, symm_symm⟩ +theorem symm_bijective : Function.Bijective (symm : (M ≃* N) → N ≃* M) := + Function.bijective_iff_has_inverse.mpr ⟨_, symm_symm, symm_symm⟩ #align mul_equiv.symm_bijective MulEquiv.symm_bijective #align add_equiv.symm_bijective AddEquiv.symm_bijective diff --git a/Mathlib/Algebra/Group/Hom/Instances.lean b/Mathlib/Algebra/Group/Hom/Instances.lean index e62c12aa40827..372a8d01165ae 100644 --- a/Mathlib/Algebra/Group/Hom/Instances.lean +++ b/Mathlib/Algebra/Group/Hom/Instances.lean @@ -345,4 +345,22 @@ def AddMonoid.End.mulRight : R →+ AddMonoid.End R := #align add_monoid.End.mul_right AddMonoid.End.mulRight #align add_monoid.End.mul_right_apply_apply AddMonoid.End.mulRight_apply_apply +lemma AddMonoid.End.mulRight_eq_mulLeft_of_commute (a : R) (h : ∀ (b : R), Commute a b) : + mulRight a = mulLeft a := + AddMonoidHom.ext fun _ ↦ (h _).eq.symm + end Semiring + +section CommSemiring + +variable {R S : Type*} [NonUnitalNonAssocCommSemiring R] + +namespace AddMonoid.End + +lemma comm_mulRight_eq_mulLeft : mulRight = (mulLeft : R →+ AddMonoid.End R) := by + ext a + exact mulRight_eq_mulLeft_of_commute _ (Commute.all _) + +end AddMonoid.End + +end CommSemiring diff --git a/Mathlib/Algebra/Group/InjSurj.lean b/Mathlib/Algebra/Group/InjSurj.lean index 32c38b9d269ac..d6ff7de4462d1 100644 --- a/Mathlib/Algebra/Group/InjSurj.lean +++ b/Mathlib/Algebra/Group/InjSurj.lean @@ -320,7 +320,7 @@ protected def addGroupWithOne {M₁} [Zero M₁] [One M₁] [Add M₁] [SMul ℕ { hf.addGroup f zero add neg sub nsmul zsmul, hf.addMonoidWithOne f zero one add nsmul nat_cast with intCast := Int.cast, - intCast_ofNat := fun n => hf (by rw [nat_cast, ←Int.cast, int_cast, Int.cast_ofNat]), + intCast_ofNat := fun n => hf (by rw [nat_cast, ← Int.cast, int_cast, Int.cast_ofNat]), intCast_negSucc := fun n => hf (by erw [int_cast, neg, nat_cast, Int.cast_negSucc] ) } #align function.injective.add_group_with_one Function.Injective.addGroupWithOne @@ -412,8 +412,10 @@ protected def monoid [Monoid M₁] (f : M₁ → M₂) (hf : Surjective f) (one (mul : ∀ x y, f (x * y) = f x * f y) (npow : ∀ (x) (n : ℕ), f (x ^ n) = f x ^ n) : Monoid M₂ := { hf.semigroup f mul, hf.mulOneClass f one mul with npow := fun n x => x ^ n, - npow_zero := hf.forall.2 fun x => by dsimp only; erw [←npow, pow_zero, ←one], - npow_succ := fun n => hf.forall.2 fun x => by dsimp only; erw [←npow, pow_succ, ←npow, ←mul] } + npow_zero := hf.forall.2 fun x => by dsimp only; erw [← npow, pow_zero, ← one], + npow_succ := fun n => hf.forall.2 fun x => by + dsimp only + erw [← npow, pow_succ, ← npow, ← mul] } #align function.surjective.monoid Function.Surjective.monoid #align function.surjective.add_monoid Function.Surjective.addMonoid diff --git a/Mathlib/Algebra/Group/Opposite.lean b/Mathlib/Algebra/Group/Opposite.lean index d6c07bb571aff..6948b4de58f1b 100644 --- a/Mathlib/Algebra/Group/Opposite.lean +++ b/Mathlib/Algebra/Group/Opposite.lean @@ -220,16 +220,16 @@ theorem op_div [DivInvMonoid α] (x y : α) : op (x / y) = (op y)⁻¹ * op x := theorem semiconjBy_op [Mul α] {a x y : α} : SemiconjBy (op a) (op y) (op x) ↔ SemiconjBy a x y := by simp only [SemiconjBy, ← op_mul, op_inj, eq_comm] #align mul_opposite.semiconj_by_op MulOpposite.semiconjBy_op -#align add_opposite.semiconj_by_op AddOpposite.semiconjBy_op +#align add_opposite.semiconj_by_op AddOpposite.addSemiconjBy_op @[to_additive (attr := simp, nolint simpComm)] theorem semiconjBy_unop [Mul α] {a x y : αᵐᵒᵖ} : SemiconjBy (unop a) (unop y) (unop x) ↔ SemiconjBy a x y := by conv_rhs => rw [← op_unop a, ← op_unop x, ← op_unop y, semiconjBy_op] #align mul_opposite.semiconj_by_unop MulOpposite.semiconjBy_unop -#align add_opposite.semiconj_by_unop AddOpposite.semiconjBy_unop +#align add_opposite.semiconj_by_unop AddOpposite.addSemiconjBy_unop -attribute [nolint simpComm] AddOpposite.semiconjBy_unop +attribute [nolint simpComm] AddOpposite.addSemiconjBy_unop @[to_additive] theorem _root_.SemiconjBy.op [Mul α] {a x y : α} (h : SemiconjBy a x y) : @@ -252,24 +252,25 @@ theorem _root_.Commute.op [Mul α] {x y : α} (h : Commute x y) : Commute (op x) #align add_commute.op AddCommute.op @[to_additive] -theorem Commute.unop [Mul α] {x y : αᵐᵒᵖ} (h : Commute x y) : Commute (unop x) (unop y) := +nonrec theorem _root_.Commute.unop [Mul α] {x y : αᵐᵒᵖ} (h : Commute x y) : + Commute (unop x) (unop y) := h.unop -#align mul_opposite.commute.unop MulOpposite.Commute.unop -#align add_opposite.commute.unop AddOpposite.Commute.unop +#align mul_opposite.commute.unop Commute.unop +#align add_opposite.commute.unop AddCommute.unop @[to_additive (attr := simp)] theorem commute_op [Mul α] {x y : α} : Commute (op x) (op y) ↔ Commute x y := semiconjBy_op #align mul_opposite.commute_op MulOpposite.commute_op -#align add_opposite.commute_op AddOpposite.commute_op +#align add_opposite.commute_op AddOpposite.addCommute_op @[to_additive (attr := simp, nolint simpComm)] theorem commute_unop [Mul α] {x y : αᵐᵒᵖ} : Commute (unop x) (unop y) ↔ Commute x y := semiconjBy_unop #align mul_opposite.commute_unop MulOpposite.commute_unop -#align add_opposite.commute_unop AddOpposite.commute_unop +#align add_opposite.commute_unop AddOpposite.addCommute_unop -attribute [nolint simpComm] AddOpposite.commute_unop +attribute [nolint simpComm] AddOpposite.addCommute_unop /-- The function `MulOpposite.op` is an additive equivalence. -/ @[simps! (config := { fullyApplied := false, simpRhs := true }) apply symm_apply] diff --git a/Mathlib/Algebra/Group/Pi.lean b/Mathlib/Algebra/Group/Pi.lean index 573046ed135e8..8cfa96ddcd5cf 100644 --- a/Mathlib/Algebra/Group/Pi.lean +++ b/Mathlib/Algebra/Group/Pi.lean @@ -563,7 +563,7 @@ theorem Pi.mulSingle_commute [∀ i, MulOneClass <| f i] : simp [hij] simp [h1, h2] #align pi.mul_single_commute Pi.mulSingle_commute -#align pi.single_commute Pi.single_commute +#align pi.single_commute Pi.single_addCommute /-- The injection into a pi group with the same values commutes. -/ @[to_additive "The injection into an additive pi group with the same values commutes."] @@ -573,7 +573,7 @@ theorem Pi.mulSingle_apply_commute [∀ i, MulOneClass <| f i] (x : ∀ i, f i) · rfl · exact Pi.mulSingle_commute hij _ _ #align pi.mul_single_apply_commute Pi.mulSingle_apply_commute -#align pi.single_apply_commute Pi.single_apply_commute +#align pi.single_apply_commute Pi.single_apply_addCommute @[to_additive] theorem Pi.update_eq_div_mul_mulSingle [∀ i, Group <| f i] (g : ∀ i : I, f i) (x : f i) : diff --git a/Mathlib/Algebra/Group/Semiconj/Defs.lean b/Mathlib/Algebra/Group/Semiconj/Defs.lean index 543c14e512999..f1e617cc8fdbe 100644 --- a/Mathlib/Algebra/Group/Semiconj/Defs.lean +++ b/Mathlib/Algebra/Group/Semiconj/Defs.lean @@ -32,7 +32,7 @@ operations (`pow_right`, field inverse etc) are in the files that define corresp variable {S M G : Type*} /-- `x` is semiconjugate to `y` by `a`, if `a * x = y * a`. -/ -@[to_additive AddSemiconjBy "`x` is additive semiconjugate to `y` by `a` if `a + x = y + a`"] +@[to_additive "`x` is additive semiconjugate to `y` by `a` if `a + x = y + a`"] def SemiconjBy [Mul M] (a x y : M) : Prop := a * x = y * a #align semiconj_by SemiconjBy @@ -59,7 +59,7 @@ theorem mul_right (h : SemiconjBy a x y) (h' : SemiconjBy a x' y') : SemiconjBy a (x * x') (y * y') := by unfold SemiconjBy -- TODO this could be done using `assoc_rw` if/when this is ported to mathlib4 - rw [←mul_assoc, h.eq, mul_assoc, h'.eq, ←mul_assoc] + rw [← mul_assoc, h.eq, mul_assoc, h'.eq, ← mul_assoc] #align semiconj_by.mul_right SemiconjBy.mul_right #align add_semiconj_by.add_right AddSemiconjBy.add_right @@ -69,7 +69,7 @@ semiconjugates `x` to `z`. -/ semiconjugates `x` to `z`."] theorem mul_left (ha : SemiconjBy a y z) (hb : SemiconjBy b x y) : SemiconjBy (a * b) x z := by unfold SemiconjBy - rw [mul_assoc, hb.eq, ←mul_assoc, ha.eq, mul_assoc] + rw [mul_assoc, hb.eq, ← mul_assoc, ha.eq, mul_assoc] #align semiconj_by.mul_left SemiconjBy.mul_left #align add_semiconj_by.add_left AddSemiconjBy.add_left @@ -143,7 +143,7 @@ end Group end SemiconjBy -@[to_additive (attr := simp) addSemiconjBy_iff_eq] +@[to_additive (attr := simp)] theorem semiconjBy_iff_eq [CancelCommMonoid M] {a x y : M} : SemiconjBy a x y ↔ x = y := ⟨fun h => mul_left_cancel (h.trans (mul_comm _ _)), fun h => by rw [h, SemiconjBy, mul_comm]⟩ #align semiconj_by_iff_eq semiconjBy_iff_eq diff --git a/Mathlib/Algebra/Group/Semiconj/Units.lean b/Mathlib/Algebra/Group/Semiconj/Units.lean index 683d5e7f233fd..065963366ee40 100644 --- a/Mathlib/Algebra/Group/Semiconj/Units.lean +++ b/Mathlib/Algebra/Group/Semiconj/Units.lean @@ -123,7 +123,7 @@ end Group end SemiconjBy /-- `a` semiconjugates `x` to `a * x * a⁻¹`. -/ -@[to_additive AddUnits.mk_addSemiconjBy "`a` semiconjugates `x` to `a + x + -a`."] +@[to_additive "`a` semiconjugates `x` to `a + x + -a`."] theorem Units.mk_semiconjBy [Monoid M] (u : Mˣ) (x : M) : SemiconjBy (↑u) x (u * x * ↑u⁻¹) := by unfold SemiconjBy; rw [Units.inv_mul_cancel_right] #align units.mk_semiconj_by Units.mk_semiconjBy diff --git a/Mathlib/Algebra/Group/TypeTags.lean b/Mathlib/Algebra/Group/TypeTags.lean index 9cf73af1c836a..fb2246c659fdc 100644 --- a/Mathlib/Algebra/Group/TypeTags.lean +++ b/Mathlib/Algebra/Group/TypeTags.lean @@ -142,13 +142,13 @@ instance [h : DecidableEq α] : DecidableEq (Multiplicative α) := h instance [h : DecidableEq α] : DecidableEq (Additive α) := h -instance instNontrivialAdditive [Nontrivial α] : Nontrivial (Additive α) := +instance Additive.instNontrivial [Nontrivial α] : Nontrivial (Additive α) := ofMul.injective.nontrivial -#align additive.nontrivial instNontrivialAdditive +#align additive.nontrivial Additive.instNontrivial -instance instNontrivialMultiplicative [Nontrivial α] : Nontrivial (Multiplicative α) := +instance Multiplicative.instNontrivial [Nontrivial α] : Nontrivial (Multiplicative α) := ofAdd.injective.nontrivial -#align multiplicative.nontrivial instNontrivialMultiplicative +#align multiplicative.nontrivial Multiplicative.instNontrivial instance Additive.add [Mul α] : Add (Additive α) where add x y := ofMul (toMul x * toMul y) diff --git a/Mathlib/Algebra/Group/UniqueProds.lean b/Mathlib/Algebra/Group/UniqueProds.lean index 92709b69811ed..704c4439e4810 100644 --- a/Mathlib/Algebra/Group/UniqueProds.lean +++ b/Mathlib/Algebra/Group/UniqueProds.lean @@ -420,7 +420,7 @@ open MulOpposite in · refine ⟨c1, hc1, d2, hd2, h12, fun c3 d3 hc3 hd3 he => ?_⟩ specialize hu ⟨_, _, ⟨_, hd1, rfl⟩, hc3, rfl⟩ ⟨_, _, hd3, ⟨_, hc2, rfl⟩, rfl⟩ rw [mul_left_cancel_iff, mul_right_cancel_iff, - mul_assoc, ←mul_assoc c3, he, mul_assoc, mul_assoc] at hu; exact hu rfl + mul_assoc, ← mul_assoc c3, he, mul_assoc, mul_assoc] at hu; exact hu rfl push_neg at h12; obtain ⟨rfl, rfl⟩ := h12 by_cases h21 : c2 ≠ 1 ∨ d1 ≠ 1 · refine ⟨c2, hc2, d1, hd1, h21, fun c4 d4 hc4 hd4 he => ?_⟩ diff --git a/Mathlib/Algebra/Group/Units.lean b/Mathlib/Algebra/Group/Units.lean index 33389c7d34461..9098001b03e7d 100644 --- a/Mathlib/Algebra/Group/Units.lean +++ b/Mathlib/Algebra/Group/Units.lean @@ -749,7 +749,7 @@ theorem val_inv_mul (h : IsUnit a) : ↑h.unit⁻¹ * a = 1 := @[to_additive (attr := simp)] theorem mul_val_inv (h : IsUnit a) : a * ↑h.unit⁻¹ = 1 := by - rw [←h.unit.mul_inv]; congr + rw [← h.unit.mul_inv]; congr #align is_unit.mul_coe_inv IsUnit.mul_val_inv #align is_add_unit.add_coe_neg IsAddUnit.add_val_neg diff --git a/Mathlib/Algebra/Group/Units/Hom.lean b/Mathlib/Algebra/Group/Units/Hom.lean index 02150f895524f..53082678279c7 100644 --- a/Mathlib/Algebra/Group/Units/Hom.lean +++ b/Mathlib/Algebra/Group/Units/Hom.lean @@ -391,7 +391,7 @@ protected theorem one_div_mul_cancel (h : IsUnit a) : 1 / a * a = 1 := by simp [ @[to_additive] theorem inv (h : IsUnit a) : IsUnit a⁻¹ := by rcases h with ⟨u, hu⟩ - rw [←hu, ← Units.val_inv_eq_inv_val] + rw [← hu, ← Units.val_inv_eq_inv_val] exact Units.isUnit _ #align is_unit.inv IsUnit.inv #align is_add_unit.neg IsAddUnit.neg diff --git a/Mathlib/Algebra/Homology/ComplexShape.lean b/Mathlib/Algebra/Homology/ComplexShape.lean index e741e56ba9d3f..af12d94e53a03 100644 --- a/Mathlib/Algebra/Homology/ComplexShape.lean +++ b/Mathlib/Algebra/Homology/ComplexShape.lean @@ -102,6 +102,10 @@ theorem symm_symm (c : ComplexShape ι) : c.symm.symm = c := by simp #align complex_shape.symm_symm ComplexShape.symm_symm +theorem symm_bijective : + Function.Bijective (ComplexShape.symm : ComplexShape ι → ComplexShape ι) := + Function.bijective_iff_has_inverse.mpr ⟨_, symm_symm, symm_symm⟩ + /-- The "composition" of two `ComplexShape`s. We need this to define "related in k steps" later. diff --git a/Mathlib/Algebra/Homology/HomologySequence.lean b/Mathlib/Algebra/Homology/HomologySequence.lean new file mode 100644 index 0000000000000..a20201f31c1c8 --- /dev/null +++ b/Mathlib/Algebra/Homology/HomologySequence.lean @@ -0,0 +1,158 @@ +/- +Copyright (c) 2023 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ +import Mathlib.Algebra.Homology.ShortComplex.HomologicalComplex +import Mathlib.Algebra.Homology.ShortComplex.SnakeLemma +import Mathlib.Algebra.Homology.ShortComplex.ShortExact +import Mathlib.Algebra.Homology.HomologicalComplexLimits + +/-! +# The homology sequence + +If `0 ⟶ X₁ ⟶ X₂ ⟶ X₃ ⟶ 0` is a short exact sequence in a category of complexes +`HomologicalComplex C c` in an abelian category (i.e. `S` is a short complex in +that category and satisfies `hS : S.ShortExact`), then whenever `i` and `j` are degrees +such that `hij : c.Rel i j`, then there is a long exact sequence : +`... ⟶ S.X₁.homology i ⟶ S.X₂.homology i ⟶ S.X₃.homology i ⟶ S.X₁.homology j ⟶ ...` (TODO). + +The proof is based on the snake lemma, similarly as it was originally done in +the Liquid Tensor Experiment. + +## References + +* https://stacks.math.columbia.edu/tag/0111 + +-/ + +open CategoryTheory Category Limits + +namespace HomologicalComplex + +section HasZeroMorphisms + +variable {C ι : Type*} [Category C] [HasZeroMorphisms C] {c : ComplexShape ι} + (K L : HomologicalComplex C c) (φ : K ⟶ L) (i j : ι) + [K.HasHomology i] [K.HasHomology j] [L.HasHomology i] [L.HasHomology j] + +/-- The morphism `K.opcycles i ⟶ K.cycles j` that is induced by `K.d i j`. -/ +noncomputable def opcyclesToCycles [K.HasHomology i] [K.HasHomology j] : + K.opcycles i ⟶ K.cycles j := + K.liftCycles (K.fromOpcycles i j) _ rfl (by simp) + +@[reassoc (attr := simp)] +lemma opcyclesToCycles_iCycles : K.opcyclesToCycles i j ≫ K.iCycles j = K.fromOpcycles i j := by + dsimp only [opcyclesToCycles] + simp + +@[reassoc] +lemma pOpcycles_opcyclesToCycles_iCycles : + K.pOpcycles i ≫ K.opcyclesToCycles i j ≫ K.iCycles j = K.d i j := by + simp [opcyclesToCycles] + +@[reassoc (attr := simp)] +lemma pOpcycles_opcyclesToCycles : + K.pOpcycles i ≫ K.opcyclesToCycles i j = K.toCycles i j := by + simp only [← cancel_mono (K.iCycles j), assoc, opcyclesToCycles_iCycles, + p_fromOpcycles, toCycles_i] + +@[reassoc (attr := simp)] +lemma homologyι_opcyclesToCycles : + K.homologyι i ≫ K.opcyclesToCycles i j = 0 := by + simp only [← cancel_mono (K.iCycles j), assoc, opcyclesToCycles_iCycles, + homologyι_comp_fromOpcycles, zero_comp] + +@[reassoc (attr := simp)] +lemma opcyclesToCycles_homologyπ : + K.opcyclesToCycles i j ≫ K.homologyπ j = 0 := by + simp only [← cancel_epi (K.pOpcycles i), + pOpcycles_opcyclesToCycles_assoc, toCycles_comp_homologyπ, comp_zero] + +variable {K L} + +@[reassoc (attr := simp)] +lemma opcyclesToCycles_naturality : + opcyclesMap φ i ≫ opcyclesToCycles L i j = opcyclesToCycles K i j ≫ cyclesMap φ j := by + simp only [← cancel_mono (L.iCycles j), ← cancel_epi (K.pOpcycles i), + assoc, p_opcyclesMap_assoc, pOpcycles_opcyclesToCycles_iCycles, Hom.comm, cyclesMap_i, + pOpcycles_opcyclesToCycles_iCycles_assoc] + +variable (C c) + +/-- The natural transformation `K.opcyclesToCycles i j : K.opcycles i ⟶ K.cycles j` for all +`K : HomologicalComplex C c`. -/ +@[simps] +noncomputable def natTransOpCyclesToCycles [CategoryWithHomology C] : + opcyclesFunctor C c i ⟶ cyclesFunctor C c j where + app K := K.opcyclesToCycles i j + +end HasZeroMorphisms + +section Preadditive + +variable {C ι : Type*} [Category C] [Preadditive C] {c : ComplexShape ι} + (K : HomologicalComplex C c) (i j : ι) (hij : c.Rel i j) + +namespace HomologySequence + +/-- The diagram `K.homology i ⟶ K.opcycles i ⟶ K.cycles j ⟶ K.homology j`. -/ +@[simp] +noncomputable def composableArrows₃ [K.HasHomology i] [K.HasHomology j] : + ComposableArrows C 3 := + ComposableArrows.mk₃ (K.homologyι i) (K.opcyclesToCycles i j) (K.homologyπ j) + +instance [K.HasHomology i] [K.HasHomology j] : + Mono ((composableArrows₃ K i j).map' 0 1) := by + dsimp + infer_instance + +instance [K.HasHomology i] [K.HasHomology j] : + Epi ((composableArrows₃ K i j).map' 2 3) := by + dsimp + infer_instance + +/-- The diagram `K.homology i ⟶ K.opcycles i ⟶ K.cycles j ⟶ K.homology j` is exact +when `c.Rel i j`. -/ +lemma composableArrows₃_exact [CategoryWithHomology C] : + (composableArrows₃ K i j).Exact := by + let S := ShortComplex.mk (K.homologyι i) (K.opcyclesToCycles i j) (by simp) + let S' := ShortComplex.mk (K.homologyι i) (K.fromOpcycles i j) (by simp) + let ι : S ⟶ S' := + { τ₁ := 𝟙 _ + τ₂ := 𝟙 _ + τ₃ := K.iCycles j } + have hS : S.Exact := by + rw [ShortComplex.exact_iff_of_epi_of_isIso_of_mono ι] + exact S'.exact_of_f_is_kernel (K.homologyIsKernel i j (c.next_eq' hij)) + let T := ShortComplex.mk (K.opcyclesToCycles i j) (K.homologyπ j) (by simp) + let T' := ShortComplex.mk (K.toCycles i j) (K.homologyπ j) (by simp) + let π : T' ⟶ T := + { τ₁ := K.pOpcycles i + τ₂ := 𝟙 _ + τ₃ := 𝟙 _ } + have hT : T.Exact := by + rw [← ShortComplex.exact_iff_of_epi_of_isIso_of_mono π] + exact T'.exact_of_g_is_cokernel (K.homologyIsCokernel i j (c.prev_eq' hij)) + apply ComposableArrows.exact_of_δ₀ + · exact hS.exact_toComposableArrows + · exact hT.exact_toComposableArrows + +variable (C) + +attribute [local simp] homologyMap_comp cyclesMap_comp opcyclesMap_comp + +/-- The functor `HomologicalComplex C c ⥤ ComposableArrows C 3` that maps `K` to the +diagram `K.homology i ⟶ K.opcycles i ⟶ K.cycles j ⟶ K.homology j`. -/ +@[simps] +noncomputable def composableArrows₃Functor [CategoryWithHomology C] : + HomologicalComplex C c ⥤ ComposableArrows C 3 where + obj K := composableArrows₃ K i j + map {K L} φ := ComposableArrows.homMk₃ (homologyMap φ i) (opcyclesMap φ i) (cyclesMap φ j) + (homologyMap φ j) (by aesop_cat) (by aesop_cat) (by aesop_cat) + +end HomologySequence + +end Preadditive + +end HomologicalComplex diff --git a/Mathlib/Algebra/Homology/ShortComplex/HomologicalComplex.lean b/Mathlib/Algebra/Homology/ShortComplex/HomologicalComplex.lean index e63f6c67a9798..a6c1637262aa5 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/HomologicalComplex.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/HomologicalComplex.lean @@ -311,6 +311,12 @@ lemma cyclesMap_i : cyclesMap φ i ≫ L.iCycles i = K.iCycles i ≫ φ.f i := lemma p_opcyclesMap : K.pOpcycles i ≫ opcyclesMap φ i = φ.f i ≫ L.pOpcycles i := ShortComplex.p_opcyclesMap _ +instance [Mono (φ.f i)] : Mono (cyclesMap φ i) := mono_of_mono_fac (cyclesMap_i φ i) + +attribute [local instance] epi_comp + +instance [Epi (φ.f i)] : Epi (opcyclesMap φ i) := epi_of_epi_fac (p_opcyclesMap φ i) + variable (K) @[simp] @@ -420,6 +426,20 @@ noncomputable def opcyclesFunctor [CategoryWithHomology C] : HomologicalComplex obj K := K.opcycles i map f := opcyclesMap f i +/-- The natural transformation `K.homologyπ i : K.cycles i ⟶ K.homology i` +for all `K : HomologicalComplex C c`. -/ +@[simps] +noncomputable def natTransHomologyπ [CategoryWithHomology C] : + cyclesFunctor C c i ⟶ homologyFunctor C c i where + app K := K.homologyπ i + +/-- The natural transformation `K.homologyι i : K.homology i ⟶ K.opcycles i` +for all `K : HomologicalComplex C c`. -/ +@[simps] +noncomputable def natTransHomologyι [CategoryWithHomology C] : + homologyFunctor C c i ⟶ opcyclesFunctor C c i where + app K := K.homologyι i + /-- The natural isomorphism `K.homology i ≅ (K.sc i).homology` for all homological complexes `K`. -/ @[simps!] @@ -436,6 +456,10 @@ noncomputable def homologyFunctorIso' [CategoryWithHomology C] shortComplexFunctor' C c i j k ⋙ ShortComplex.homologyFunctor C := homologyFunctorIso C c j ≪≫ isoWhiskerRight (natIsoSc' C c i j k hi hk) _ +instance [CategoryWithHomology C] : (homologyFunctor C c i).PreservesZeroMorphisms where +instance [CategoryWithHomology C] : (opcyclesFunctor C c i).PreservesZeroMorphisms where +instance [CategoryWithHomology C] : (cyclesFunctor C c i).PreservesZeroMorphisms where + end end diff --git a/Mathlib/Algebra/Homology/ShortComplex/SnakeLemma.lean b/Mathlib/Algebra/Homology/ShortComplex/SnakeLemma.lean index 6e7d3cf947350..0b78caea98716 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/SnakeLemma.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/SnakeLemma.lean @@ -223,6 +223,11 @@ noncomputable def P := pullback S.L₁.g S.v₀₁.τ₃ /-- The canonical map `P ⟶ L₂.X₂`. -/ noncomputable def φ₂ : S.P ⟶ S.L₂.X₂ := pullback.fst ≫ S.v₁₂.τ₂ +@[reassoc (attr := simp)] +lemma lift_φ₂ {A : C} (a : A ⟶ S.L₁.X₂) (b : A ⟶ S.L₀.X₃) (h : a ≫ S.L₁.g = b ≫ S.v₀₁.τ₃) : + pullback.lift a b h ≫ S.φ₂ = a ≫ S.v₁₂.τ₂ := by + simp [φ₂] + /-- The canonical map `P ⟶ L₂.X₁`. -/ noncomputable def φ₁ : S.P ⟶ S.L₂.X₁ := S.L₂_exact.lift S.φ₂ @@ -362,6 +367,15 @@ lemma snake_lemma : S.composableArrows.Exact := (exact_of_δ₀ S.L₂'_exact.exact_toComposableArrows S.L₃_exact.exact_toComposableArrows)) +lemma δ_eq {A : C} (x₃ : A ⟶ S.L₀.X₃) (x₂ : A ⟶ S.L₁.X₂) (x₁ : A ⟶ S.L₂.X₁) + (h₂ : x₂ ≫ S.L₁.g = x₃ ≫ S.v₀₁.τ₃) (h₁ : x₁ ≫ S.L₂.f = x₂ ≫ S.v₁₂.τ₂) : + x₃ ≫ S.δ = x₁ ≫ S.v₂₃.τ₁ := by + have H := (pullback.lift x₂ x₃ h₂) ≫= S.snd_δ + rw [pullback.lift_snd_assoc] at H + rw [H, ← assoc] + congr 1 + simp only [← cancel_mono S.L₂.f, assoc, φ₁_L₂_f, lift_φ₂, h₁] + variable (S₁ S₂ S₃ : SnakeInput C) /-- A morphism of snake inputs involve four morphisms of short complexes diff --git a/Mathlib/Algebra/IndicatorFunction.lean b/Mathlib/Algebra/IndicatorFunction.lean index 1853feb551b2d..ad96c463add33 100644 --- a/Mathlib/Algebra/IndicatorFunction.lean +++ b/Mathlib/Algebra/IndicatorFunction.lean @@ -662,7 +662,7 @@ theorem _root_.Finset.prod_mulIndicator_eq_prod_filter (s : Finset ι) (f : ι @[to_additive] theorem mulIndicator_finset_prod (I : Finset ι) (s : Set α) (f : ι → α → M) : mulIndicator s (∏ i in I, f i) = ∏ i in I, mulIndicator s (f i) := - (mulIndicatorHom M s).map_prod _ _ + map_prod (mulIndicatorHom M s) _ _ #align set.mul_indicator_finset_prod Set.mulIndicator_finset_prod #align set.indicator_finset_sum Set.indicator_finset_sum diff --git a/Mathlib/Algebra/IsPrimePow.lean b/Mathlib/Algebra/IsPrimePow.lean index 72ac069276671..e3f8d507fdcf8 100644 --- a/Mathlib/Algebra/IsPrimePow.lean +++ b/Mathlib/Algebra/IsPrimePow.lean @@ -87,7 +87,7 @@ theorem isPrimePow_nat_iff_bounded (n : ℕ) : refine' Iff.symm ⟨fun ⟨p, _, k, _, hp, hk, hn⟩ => ⟨p, k, hp, hk, hn⟩, _⟩ rintro ⟨p, k, hp, hk, rfl⟩ refine' ⟨p, _, k, (Nat.lt_pow_self hp.one_lt _).le, hp, hk, rfl⟩ - conv => { lhs; rw [←(pow_one p)] } + conv => { lhs; rw [← (pow_one p)] } exact (Nat.pow_le_iff_le_right hp.two_le).mpr hk #align is_prime_pow_nat_iff_bounded isPrimePow_nat_iff_bounded diff --git a/Mathlib/Algebra/Jordan/Basic.lean b/Mathlib/Algebra/Jordan/Basic.lean index 10b66852e895f..12b6b252a178f 100644 --- a/Mathlib/Algebra/Jordan/Basic.lean +++ b/Mathlib/Algebra/Jordan/Basic.lean @@ -86,24 +86,23 @@ class IsJordan [Mul A] : Prop where #align is_jordan IsJordan /-- A commutative Jordan multipication -/ -class IsCommJordan [Mul A] : Prop where - mul_comm : ∀ a b : A, a * b = b * a +class IsCommJordan [CommMagma A] : Prop where lmul_comm_rmul_rmul : ∀ a b : A, a * b * (a * a) = a * (b * (a * a)) #align is_comm_jordan IsCommJordan -- see Note [lower instance priority] /-- A (commutative) Jordan multiplication is also a Jordan multipication -/ -instance (priority := 100) IsCommJordan.toIsJordan [Mul A] [IsCommJordan A] : IsJordan A where - lmul_comm_rmul a b := by rw [IsCommJordan.mul_comm, IsCommJordan.mul_comm a b] +instance (priority := 100) IsCommJordan.toIsJordan [CommMagma A] [IsCommJordan A] : IsJordan A where + lmul_comm_rmul a b := by rw [mul_comm, mul_comm a b] lmul_lmul_comm_lmul a b := by - rw [IsCommJordan.mul_comm (a * a) (a * b), IsCommJordan.lmul_comm_rmul_rmul, - IsCommJordan.mul_comm b (a * a)] + rw [mul_comm (a * a) (a * b), IsCommJordan.lmul_comm_rmul_rmul, + mul_comm b (a * a)] lmul_comm_rmul_rmul := IsCommJordan.lmul_comm_rmul_rmul lmul_lmul_comm_rmul a b := by - rw [IsCommJordan.mul_comm (a * a) (b * a), IsCommJordan.mul_comm b a, - IsCommJordan.lmul_comm_rmul_rmul, IsCommJordan.mul_comm, IsCommJordan.mul_comm b (a * a)] + rw [mul_comm (a * a) (b * a), mul_comm b a, + IsCommJordan.lmul_comm_rmul_rmul, mul_comm, mul_comm b (a * a)] rmul_comm_rmul_rmul a b := by - rw [IsCommJordan.mul_comm b a, IsCommJordan.lmul_comm_rmul_rmul, IsCommJordan.mul_comm] + rw [mul_comm b a, IsCommJordan.lmul_comm_rmul_rmul, mul_comm] #align is_comm_jordan.to_is_jordan IsCommJordan.toIsJordan -- see Note [lower instance priority] @@ -118,7 +117,6 @@ instance (priority := 100) Semigroup.isJordan [Semigroup A] : IsJordan A where -- see Note [lower instance priority] instance (priority := 100) CommSemigroup.isCommJordan [CommSemigroup A] : IsCommJordan A where - mul_comm := mul_comm lmul_comm_rmul_rmul _ _ := mul_assoc _ _ _ #align comm_semigroup.is_comm_jordan CommSemigroup.isCommJordan @@ -162,7 +160,7 @@ theorem commute_rmul_rmul_sq (a : A) : Commute (R a) (R (a * a)) := end Commute -variable {A} [NonUnitalNonAssocRing A] [IsCommJordan A] +variable {A} [NonUnitalNonAssocCommRing A] [IsCommJordan A] /-! The endomorphisms on an additive monoid `AddMonoid.End` form a `Ring`, and this may be equipped @@ -175,7 +173,7 @@ theorem two_nsmul_lie_lmul_lmul_add_eq_lie_lmul_lmul_add (a b : A) : suffices 2 • ⁅L a, L (a * b)⁆ + 2 • ⁅L b, L (b * a)⁆ + ⁅L b, L (a * a)⁆ + ⁅L a, L (b * b)⁆ = 0 by rwa [← sub_eq_zero, ← sub_sub, sub_eq_add_neg, sub_eq_add_neg, lie_skew, lie_skew, nsmul_add] convert (commute_lmul_lmul_sq (a + b)).lie_eq using 1 - simp only [add_mul, mul_add, map_add, lie_add, add_lie, IsCommJordan.mul_comm b a, + simp only [add_mul, mul_add, map_add, lie_add, add_lie, mul_comm b a, (commute_lmul_lmul_sq a).lie_eq, (commute_lmul_lmul_sq b).lie_eq, zero_add, add_zero, two_smul] abel #align two_nsmul_lie_lmul_lmul_add_eq_lie_lmul_lmul_add two_nsmul_lie_lmul_lmul_add_eq_lie_lmul_lmul_add @@ -189,7 +187,7 @@ private theorem aux0 {a b c : A} : ⁅L (a + b + c), L ((a + b + c) * (a + b + c rw [add_mul, add_mul] iterate 6 rw [mul_add] iterate 10 rw [map_add] - rw [IsCommJordan.mul_comm b a, IsCommJordan.mul_comm c a, IsCommJordan.mul_comm c b] + rw [mul_comm b a, mul_comm c a, mul_comm c b] iterate 3 rw [two_smul] simp only [lie_add, add_lie, commute_lmul_lmul_sq, zero_add, add_zero] abel @@ -235,9 +233,9 @@ private theorem aux3 {a b c : A} : 2 • ⁅L a, L (b * c)⁆ + 2 • ⁅L b, L (c * a)⁆ + 2 • ⁅L c, L (a * b)⁆ := by rw [add_left_eq_self] -- Porting note: was `nth_rw` instead of `conv_lhs` - conv_lhs => enter [1, 1, 2, 2, 2]; rw [IsCommJordan.mul_comm a b] - conv_lhs => enter [1, 2, 2, 2, 1]; rw [IsCommJordan.mul_comm c a] - conv_lhs => enter [ 2, 2, 2, 2]; rw [IsCommJordan.mul_comm b c] + conv_lhs => enter [1, 1, 2, 2, 2]; rw [mul_comm a b] + conv_lhs => enter [1, 2, 2, 2, 1]; rw [mul_comm c a] + conv_lhs => enter [ 2, 2, 2, 2]; rw [mul_comm b c] iterate 3 rw [two_nsmul_lie_lmul_lmul_add_eq_lie_lmul_lmul_add] iterate 2 rw [← lie_skew (L (a * a)), ← lie_skew (L (b * b)), ← lie_skew (L (c * c))] abel diff --git a/Mathlib/Algebra/Lie/Basic.lean b/Mathlib/Algebra/Lie/Basic.lean index 441ab5d2e19a5..c6ebe8c39e216 100644 --- a/Mathlib/Algebra/Lie/Basic.lean +++ b/Mathlib/Algebra/Lie/Basic.lean @@ -633,6 +633,9 @@ theorem symm_symm (e : L₁ ≃ₗ⁅R⁆ L₂) : e.symm.symm = e := by rfl #align lie_equiv.symm_symm LieEquiv.symm_symm +theorem symm_bijective : Function.Bijective (LieEquiv.symm : (L₁ ≃ₗ⁅R⁆ L₂) → L₂ ≃ₗ⁅R⁆ L₁) := + Function.bijective_iff_has_inverse.mpr ⟨_, symm_symm, symm_symm⟩ + @[simp] theorem apply_symm_apply (e : L₁ ≃ₗ⁅R⁆ L₂) : ∀ x, e (e.symm x) = x := e.toLinearEquiv.apply_symm_apply @@ -1108,6 +1111,10 @@ theorem symm_symm (e : M ≃ₗ⁅R,L⁆ N) : e.symm.symm = e := by rfl #align lie_module_equiv.symm_symm LieModuleEquiv.symm_symm +theorem symm_bijective : + Function.Bijective (LieModuleEquiv.symm : (M ≃ₗ⁅R,L⁆ N) → N ≃ₗ⁅R,L⁆ M) := + Function.bijective_iff_has_inverse.mpr ⟨_, symm_symm, symm_symm⟩ + /-- Lie module equivalences are transitive. -/ @[trans] def trans (e₁ : M ≃ₗ⁅R,L⁆ N) (e₂ : N ≃ₗ⁅R,L⁆ P) : M ≃ₗ⁅R,L⁆ P := diff --git a/Mathlib/Algebra/Lie/Weights/Basic.lean b/Mathlib/Algebra/Lie/Weights/Basic.lean index fd7d321fbfb4c..a2c30572e29f0 100644 --- a/Mathlib/Algebra/Lie/Weights/Basic.lean +++ b/Mathlib/Algebra/Lie/Weights/Basic.lean @@ -455,7 +455,8 @@ lemma posFittingComp_map_incl_sup_of_codisjoint [IsNoetherian R M] [IsArtinian R (eventually_iInf_lowerCentralSeries_eq R L M) obtain ⟨hl₁, hl₂, hl₃⟩ := hl l (le_refl _) simp_rw [← iInf_lowerCentralSeries_eq_posFittingComp, hl₁, hl₂, hl₃, - LieSubmodule.lowerCentralSeries_map_eq_lcs, ←LieSubmodule.lcs_sup, lowerCentralSeries, h.eq_top] + LieSubmodule.lowerCentralSeries_map_eq_lcs, ← LieSubmodule.lcs_sup, lowerCentralSeries, + h.eq_top] lemma weightSpace_weightSpaceOf_map_incl (x : L) (χ : L → R) : (weightSpace (weightSpaceOf M (χ x) x) χ).map (weightSpaceOf M (χ x) x).incl = @@ -661,7 +662,8 @@ lemma iSup_weightSpace_eq_top [IsTriangularizable K L M] : rw [← hχ, weightSpace_weightSpaceOf_map_incl] simp_rw [biSup_congr this, ← LieSubmodule.map_iSup, ih, LieModuleHom.map_top, LieSubmodule.range_incl] - simpa only [←ih, iSup_comm (ι := K), iSup_iSup_eq_right] using iSup_weightSpaceOf_eq_top K L M y + simpa only [← ih, iSup_comm (ι := K), iSup_iSup_eq_right] using + iSup_weightSpaceOf_eq_top K L M y end field diff --git a/Mathlib/Algebra/Module/Basic.lean b/Mathlib/Algebra/Module/Basic.lean index 91f19844a5a7d..122a579dd7ff8 100644 --- a/Mathlib/Algebra/Module/Basic.lean +++ b/Mathlib/Algebra/Module/Basic.lean @@ -319,7 +319,7 @@ def Module.addCommMonoidToAddCommGroup [Ring R] [AddCommMonoid M] [Module R M] : zsmul := fun z a => (z : R) • a zsmul_zero' := fun a => by simpa only [Int.cast_zero] using zero_smul R a zsmul_succ' := fun z a => by simp [add_comm, add_smul] - zsmul_neg' := fun z a => by simp [←smul_assoc, neg_one_smul] } + zsmul_neg' := fun z a => by simp [← smul_assoc, neg_one_smul] } #align module.add_comm_monoid_to_add_comm_group Module.addCommMonoidToAddCommGroup variable {R} diff --git a/Mathlib/Algebra/Module/BigOperators.lean b/Mathlib/Algebra/Module/BigOperators.lean index 9c9636734e65b..8ad17e66616dc 100644 --- a/Mathlib/Algebra/Module/BigOperators.lean +++ b/Mathlib/Algebra/Module/BigOperators.lean @@ -36,7 +36,7 @@ theorem Multiset.sum_smul_sum {s : Multiset R} {t : Multiset M} : #align multiset.sum_smul_sum Multiset.sum_smul_sum theorem Finset.sum_smul {f : ι → R} {s : Finset ι} {x : M} : - (∑ i in s, f i) • x = ∑ i in s, f i • x := ((smulAddHom R M).flip x).map_sum f s + (∑ i in s, f i) • x = ∑ i in s, f i • x := map_sum ((smulAddHom R M).flip x) f s #align finset.sum_smul Finset.sum_smul theorem Finset.sum_smul_sum {f : α → R} {g : β → M} {s : Finset α} {t : Finset β} : diff --git a/Mathlib/Algebra/Module/Equiv.lean b/Mathlib/Algebra/Module/Equiv.lean index ab957e15c7de6..03c3042fb8c9c 100644 --- a/Mathlib/Algebra/Module/Equiv.lean +++ b/Mathlib/Algebra/Module/Equiv.lean @@ -519,9 +519,7 @@ theorem symm_symm (e : M ≃ₛₗ[σ] M₂) : e.symm.symm = e := by theorem symm_bijective [Module R M] [Module S M₂] [RingHomInvPair σ' σ] [RingHomInvPair σ σ'] : Function.Bijective (symm : (M ≃ₛₗ[σ] M₂) → M₂ ≃ₛₗ[σ'] M) := - Equiv.bijective - ⟨(symm : (M ≃ₛₗ[σ] M₂) → M₂ ≃ₛₗ[σ'] M), (symm : (M₂ ≃ₛₗ[σ'] M) → M ≃ₛₗ[σ] M₂), symm_symm, - symm_symm⟩ + Function.bijective_iff_has_inverse.mpr ⟨_, symm_symm, symm_symm⟩ #align linear_equiv.symm_bijective LinearEquiv.symm_bijective @[simp] diff --git a/Mathlib/Algebra/Module/LocalizedModule.lean b/Mathlib/Algebra/Module/LocalizedModule.lean index bc9d0e115a573..ceca2a833a2fd 100644 --- a/Mathlib/Algebra/Module/LocalizedModule.lean +++ b/Mathlib/Algebra/Module/LocalizedModule.lean @@ -717,8 +717,8 @@ noncomputable def fromLocalizedModule' : LocalizedModule S M → M' := fun p => rintro ⟨a, b⟩ ⟨a', b'⟩ ⟨c, eq1⟩ dsimp -- Porting note: We remove `generalize_proofs h1 h2`. - erw [Module.End_algebraMap_isUnit_inv_apply_eq_iff, ←map_smul, ←map_smul, - Module.End_algebraMap_isUnit_inv_apply_eq_iff', ←map_smul] + erw [Module.End_algebraMap_isUnit_inv_apply_eq_iff, ← map_smul, ← map_smul, + Module.End_algebraMap_isUnit_inv_apply_eq_iff', ← map_smul] exact (IsLocalizedModule.eq_iff_exists S f).mpr ⟨c, eq1.symm⟩) #align is_localized_module.from_localized_module' IsLocalizedModule.fromLocalizedModule' @@ -736,8 +736,8 @@ theorem fromLocalizedModule'_add (x y : LocalizedModule S M) : intro a a' b b' simp only [LocalizedModule.mk_add_mk, fromLocalizedModule'_mk] -- Porting note: We remove `generalize_proofs h1 h2 h3`. - erw [Module.End_algebraMap_isUnit_inv_apply_eq_iff, smul_add, ←map_smul, ←map_smul, - ←map_smul, map_add] + erw [Module.End_algebraMap_isUnit_inv_apply_eq_iff, smul_add, ← map_smul, ← map_smul, + ← map_smul, map_add] congr 1 all_goals rw [Module.End_algebraMap_isUnit_inv_apply_eq_iff'] · erw [mul_smul, f.map_smul] diff --git a/Mathlib/Algebra/MonoidAlgebra/Ideal.lean b/Mathlib/Algebra/MonoidAlgebra/Ideal.lean index 89f7d39028634..3bbe854b730d6 100644 --- a/Mathlib/Algebra/MonoidAlgebra/Ideal.lean +++ b/Mathlib/Algebra/MonoidAlgebra/Ideal.lean @@ -44,7 +44,7 @@ theorem MonoidAlgebra.mem_ideal_span_of_image [Monoid G] [Semiring k] {s : Set G change _ ↔ x ∈ RHS constructor · revert x - rw [←SetLike.le_def] -- porting note: refine needs this even though it's defeq? + rw [← SetLike.le_def] -- porting note: refine needs this even though it's defeq? refine Ideal.span_le.2 ?_ rintro _ ⟨i, hi, rfl⟩ m hm refine' ⟨_, hi, 1, _⟩ diff --git a/Mathlib/Algebra/Order/CompleteField.lean b/Mathlib/Algebra/Order/CompleteField.lean index a7570d25245e5..453f2cae49456 100644 --- a/Mathlib/Algebra/Order/CompleteField.lean +++ b/Mathlib/Algebra/Order/CompleteField.lean @@ -66,7 +66,7 @@ instance (priority := 100) ConditionallyCompleteLinearOrderedField.to_archimedea [ConditionallyCompleteLinearOrderedField α] : Archimedean α := archimedean_iff_nat_lt.2 (by - by_contra' h + by_contra! h obtain ⟨x, h⟩ := h have := csSup_le _ _ (range_nonempty Nat.cast) (forall_range_iff.2 fun m => diff --git a/Mathlib/Algebra/Order/Field/Basic.lean b/Mathlib/Algebra/Order/Field/Basic.lean index 45c19a818fb23..22490b49341c4 100644 --- a/Mathlib/Algebra/Order/Field/Basic.lean +++ b/Mathlib/Algebra/Order/Field/Basic.lean @@ -922,11 +922,11 @@ theorem sub_one_div_inv_le_two (a2 : 2 ≤ a) : (1 - 1 / a)⁻¹ ≤ 2 := by -- Take inverses on both sides to obtain `2⁻¹ ≤ 1 - 1 / a` refine' (inv_le_inv_of_le (inv_pos.2 <| zero_lt_two' α) _).trans_eq (inv_inv (2 : α)) -- move `1 / a` to the left and `2⁻¹` to the right. - rw [le_sub_iff_add_le, add_comm, ←le_sub_iff_add_le] + rw [le_sub_iff_add_le, add_comm, ← le_sub_iff_add_le] -- take inverses on both sides and use the assumption `2 ≤ a`. convert (one_div a).le.trans (inv_le_inv_of_le zero_lt_two a2) using 1 -- show `1 - 1 / 2 = 1 / 2`. - rw [sub_eq_iff_eq_add, ←two_mul, mul_inv_cancel two_ne_zero] + rw [sub_eq_iff_eq_add, ← two_mul, mul_inv_cancel two_ne_zero] #align sub_one_div_inv_le_two sub_one_div_inv_le_two /-! ### Results about `IsLUB` -/ diff --git a/Mathlib/Algebra/Order/Floor.lean b/Mathlib/Algebra/Order/Floor.lean index 13059e11e6261..d82668adfe432 100644 --- a/Mathlib/Algebra/Order/Floor.lean +++ b/Mathlib/Algebra/Order/Floor.lean @@ -453,11 +453,12 @@ theorem floor_add_nat (ha : 0 ≤ a) (n : ℕ) : ⌊a + n⌋₊ = ⌊a⌋₊ + n theorem floor_add_one (ha : 0 ≤ a) : ⌊a + 1⌋₊ = ⌊a⌋₊ + 1 := by -- Porting note: broken `convert floor_add_nat ha 1` - rw [←cast_one, floor_add_nat ha 1] + rw [← cast_one, floor_add_nat ha 1] #align nat.floor_add_one Nat.floor_add_one +-- See note [no_index around OfNat.ofNat] theorem floor_add_ofNat (ha : 0 ≤ a) (n : ℕ) [n.AtLeastTwo] : - ⌊a + OfNat.ofNat n⌋₊ = ⌊a⌋₊ + OfNat.ofNat n := + ⌊a + (no_index (OfNat.ofNat n))⌋₊ = ⌊a⌋₊ + OfNat.ofNat n := floor_add_nat ha n @[simp] @@ -476,9 +477,10 @@ theorem floor_sub_nat [Sub α] [OrderedSub α] [ExistsAddOfLE α] (a : α) (n : theorem floor_sub_one [Sub α] [OrderedSub α] [ExistsAddOfLE α] (a : α) : ⌊a - 1⌋₊ = ⌊a⌋₊ - 1 := mod_cast floor_sub_nat a 1 +-- See note [no_index around OfNat.ofNat] @[simp] theorem floor_sub_ofNat [Sub α] [OrderedSub α] [ExistsAddOfLE α] (a : α) (n : ℕ) [n.AtLeastTwo] : - ⌊a - OfNat.ofNat n⌋₊ = ⌊a⌋₊ - OfNat.ofNat n := + ⌊a - (no_index (OfNat.ofNat n))⌋₊ = ⌊a⌋₊ - OfNat.ofNat n := floor_sub_nat a n theorem ceil_add_nat (ha : 0 ≤ a) (n : ℕ) : ⌈a + n⌉₊ = ⌈a⌉₊ + n := @@ -496,8 +498,9 @@ theorem ceil_add_one (ha : 0 ≤ a) : ⌈a + 1⌉₊ = ⌈a⌉₊ + 1 := by rw [cast_one.symm, ceil_add_nat ha 1] #align nat.ceil_add_one Nat.ceil_add_one +-- See note [no_index around OfNat.ofNat] theorem ceil_add_ofNat (ha : 0 ≤ a) (n : ℕ) [n.AtLeastTwo] : - ⌈a + OfNat.ofNat n⌉₊ = ⌈a⌉₊ + OfNat.ofNat n := + ⌈a + (no_index (OfNat.ofNat n))⌉₊ = ⌈a⌉₊ + OfNat.ofNat n := ceil_add_nat ha n theorem ceil_lt_add_one (ha : 0 ≤ a) : (⌈a⌉₊ : α) < a + 1 := @@ -543,8 +546,9 @@ theorem floor_div_nat (a : α) (n : ℕ) : ⌊a / n⌋₊ = ⌊a⌋₊ / n := by · exact cast_pos.2 hn #align nat.floor_div_nat Nat.floor_div_nat +-- See note [no_index around OfNat.ofNat] theorem floor_div_ofNat (a : α) (n : ℕ) [n.AtLeastTwo] : - ⌊a / OfNat.ofNat n⌋₊ = ⌊a⌋₊ / OfNat.ofNat n := + ⌊a / (no_index (OfNat.ofNat n))⌋₊ = ⌊a⌋₊ / OfNat.ofNat n := floor_div_nat a n /-- Natural division is the floor of field division. -/ @@ -738,7 +742,9 @@ theorem floor_zero : ⌊(0 : α)⌋ = 0 := by rw [← cast_zero, floor_intCast] theorem floor_one : ⌊(1 : α)⌋ = 1 := by rw [← cast_one, floor_intCast] #align int.floor_one Int.floor_one -@[simp] theorem floor_ofNat (n : ℕ) [n.AtLeastTwo] : ⌊(OfNat.ofNat n : α)⌋ = n := floor_natCast n +-- See note [no_index around OfNat.ofNat] +@[simp] theorem floor_ofNat (n : ℕ) [n.AtLeastTwo] : ⌊(no_index (OfNat.ofNat n : α))⌋ = n := + floor_natCast n @[mono] theorem floor_mono : Monotone (floor : α → ℤ) := @@ -786,9 +792,10 @@ theorem floor_int_add (z : ℤ) (a : α) : ⌊↑z + a⌋ = z + ⌊a⌋ := by theorem floor_add_nat (a : α) (n : ℕ) : ⌊a + n⌋ = ⌊a⌋ + n := by rw [← Int.cast_ofNat, floor_add_int] #align int.floor_add_nat Int.floor_add_nat +-- See note [no_index around OfNat.ofNat] @[simp] theorem floor_add_ofNat (a : α) (n : ℕ) [n.AtLeastTwo] : - ⌊a + OfNat.ofNat n⌋ = ⌊a⌋ + OfNat.ofNat n := + ⌊a + (no_index (OfNat.ofNat n))⌋ = ⌊a⌋ + OfNat.ofNat n := floor_add_nat a n @[simp] @@ -796,9 +803,10 @@ theorem floor_nat_add (n : ℕ) (a : α) : ⌊↑n + a⌋ = n + ⌊a⌋ := by rw [← Int.cast_ofNat, floor_int_add] #align int.floor_nat_add Int.floor_nat_add +-- See note [no_index around OfNat.ofNat] @[simp] theorem floor_ofNat_add (n : ℕ) [n.AtLeastTwo] (a : α) : - ⌊OfNat.ofNat n + a⌋ = OfNat.ofNat n + ⌊a⌋ := + ⌊(no_index (OfNat.ofNat n)) + a⌋ = OfNat.ofNat n + ⌊a⌋ := floor_nat_add n a @[simp] @@ -812,9 +820,10 @@ theorem floor_sub_nat (a : α) (n : ℕ) : ⌊a - n⌋ = ⌊a⌋ - n := by rw [ @[simp] theorem floor_sub_one (a : α) : ⌊a - 1⌋ = ⌊a⌋ - 1 := mod_cast floor_sub_nat a 1 +-- See note [no_index around OfNat.ofNat] @[simp] theorem floor_sub_ofNat (a : α) (n : ℕ) [n.AtLeastTwo] : - ⌊a - OfNat.ofNat n⌋ = ⌊a⌋ - OfNat.ofNat n := + ⌊a - (no_index (OfNat.ofNat n))⌋ = ⌊a⌋ - OfNat.ofNat n := floor_sub_nat a n theorem abs_sub_lt_one_of_floor_eq_floor {α : Type*} [LinearOrderedCommRing α] [FloorRing α] @@ -883,8 +892,10 @@ theorem fract_add_nat (a : α) (m : ℕ) : fract (a + m) = fract a := by @[simp] theorem fract_add_one (a : α) : fract (a + 1) = fract a := mod_cast fract_add_nat a 1 +-- See note [no_index around OfNat.ofNat] @[simp] -theorem fract_add_ofNat (a : α) (n : ℕ) [n.AtLeastTwo] : fract (a + OfNat.ofNat n) = fract a := +theorem fract_add_ofNat (a : α) (n : ℕ) [n.AtLeastTwo] : + fract (a + (no_index (OfNat.ofNat n))) = fract a := fract_add_nat a n @[simp] @@ -897,8 +908,10 @@ theorem fract_nat_add (n : ℕ) (a : α) : fract (↑n + a) = fract a := by rw [ @[simp] theorem fract_one_add (a : α) : fract (1 + a) = fract a := mod_cast fract_nat_add 1 a +-- See note [no_index around OfNat.ofNat] @[simp] -theorem fract_ofNat_add (n : ℕ) [n.AtLeastTwo] (a : α) : fract (OfNat.ofNat n + a) = fract a := +theorem fract_ofNat_add (n : ℕ) [n.AtLeastTwo] (a : α) : + fract ((no_index (OfNat.ofNat n)) + a) = fract a := fract_nat_add n a @[simp] @@ -916,8 +929,10 @@ theorem fract_sub_nat (a : α) (n : ℕ) : fract (a - n) = fract a := by @[simp] theorem fract_sub_one (a : α) : fract (a - 1) = fract a := mod_cast fract_sub_nat a 1 +-- See note [no_index around OfNat.ofNat] @[simp] -theorem fract_sub_ofNat (a : α) (n : ℕ) [n.AtLeastTwo] : fract (a - OfNat.ofNat n) = fract a := +theorem fract_sub_ofNat (a : α) (n : ℕ) [n.AtLeastTwo] : + fract (a - (no_index (OfNat.ofNat n))) = fract a := fract_sub_nat a n -- Was a duplicate lemma under a bad name @@ -985,8 +1000,11 @@ theorem fract_intCast (z : ℤ) : fract (z : α) = 0 := by theorem fract_natCast (n : ℕ) : fract (n : α) = 0 := by simp [fract] #align int.fract_nat_cast Int.fract_natCast +-- See note [no_index around OfNat.ofNat] @[simp] -theorem fract_ofNat (n : ℕ) [n.AtLeastTwo] : fract (OfNat.ofNat n : α) = 0 := fract_natCast n +theorem fract_ofNat (n : ℕ) [n.AtLeastTwo] : + fract ((no_index (OfNat.ofNat n)) : α) = 0 := + fract_natCast n -- porting note: simp can prove this -- @[simp] @@ -1213,8 +1231,9 @@ theorem ceil_natCast (n : ℕ) : ⌈(n : α)⌉ = n := eq_of_forall_ge_iff fun a => by rw [ceil_le, ← cast_ofNat, cast_le] #align int.ceil_nat_cast Int.ceil_natCast +-- See note [no_index around OfNat.ofNat] @[simp] -theorem ceil_ofNat (n : ℕ) [n.AtLeastTwo] : ⌈(OfNat.ofNat n : α)⌉ = n := ceil_natCast n +theorem ceil_ofNat (n : ℕ) [n.AtLeastTwo] : ⌈(no_index (OfNat.ofNat n : α))⌉ = n := ceil_natCast n theorem ceil_mono : Monotone (ceil : α → ℤ) := gc_ceil_coe.monotone_l @@ -1235,11 +1254,13 @@ theorem ceil_add_nat (a : α) (n : ℕ) : ⌈a + n⌉ = ⌈a⌉ + n := by rw [ @[simp] theorem ceil_add_one (a : α) : ⌈a + 1⌉ = ⌈a⌉ + 1 := by -- Porting note: broken `convert ceil_add_int a (1 : ℤ)` - rw [←ceil_add_int a (1 : ℤ), cast_one] + rw [← ceil_add_int a (1 : ℤ), cast_one] #align int.ceil_add_one Int.ceil_add_one +-- See note [no_index around OfNat.ofNat] @[simp] -theorem ceil_add_ofNat (a : α) (n : ℕ) [n.AtLeastTwo] : ⌈a + OfNat.ofNat n⌉ = ⌈a⌉ + OfNat.ofNat n := +theorem ceil_add_ofNat (a : α) (n : ℕ) [n.AtLeastTwo] : + ⌈a + (no_index (OfNat.ofNat n))⌉ = ⌈a⌉ + OfNat.ofNat n := ceil_add_nat a n @[simp] @@ -1258,8 +1279,10 @@ theorem ceil_sub_one (a : α) : ⌈a - 1⌉ = ⌈a⌉ - 1 := by rw [eq_sub_iff_add_eq, ← ceil_add_one, sub_add_cancel] #align int.ceil_sub_one Int.ceil_sub_one +-- See note [no_index around OfNat.ofNat] @[simp] -theorem ceil_sub_ofNat (a : α) (n : ℕ) [n.AtLeastTwo] : ⌈a - OfNat.ofNat n⌉ = ⌈a⌉ - OfNat.ofNat n := +theorem ceil_sub_ofNat (a : α) (n : ℕ) [n.AtLeastTwo] : + ⌈a - (no_index (OfNat.ofNat n))⌉ = ⌈a⌉ - OfNat.ofNat n := ceil_sub_nat a n theorem ceil_lt_add_one (a : α) : (⌈a⌉ : α) < a + 1 := by @@ -1438,8 +1461,10 @@ theorem round_one : round (1 : α) = 1 := by simp [round] theorem round_natCast (n : ℕ) : round (n : α) = n := by simp [round] #align round_nat_cast round_natCast +-- See note [no_index around OfNat.ofNat] @[simp] -theorem round_ofNat (n : ℕ) [n.AtLeastTwo] : round (OfNat.ofNat n : α) = n := round_natCast n +theorem round_ofNat (n : ℕ) [n.AtLeastTwo] : round (no_index (OfNat.ofNat n : α)) = n := + round_natCast n @[simp] theorem round_intCast (n : ℤ) : round (n : α) = n := by simp [round] @@ -1453,7 +1478,7 @@ theorem round_add_int (x : α) (y : ℤ) : round (x + y) = round x + y := by @[simp] theorem round_add_one (a : α) : round (a + 1) = round a + 1 := by -- Porting note: broken `convert round_add_int a 1` - rw [←round_add_int a 1, cast_one] + rw [← round_add_int a 1, cast_one] #align round_add_one round_add_one @[simp] @@ -1466,7 +1491,7 @@ theorem round_sub_int (x : α) (y : ℤ) : round (x - y) = round x - y := by @[simp] theorem round_sub_one (a : α) : round (a - 1) = round a - 1 := by -- Porting note: broken `convert round_sub_int a 1` - rw [←round_sub_int a 1, cast_one] + rw [← round_sub_int a 1, cast_one] #align round_sub_one round_sub_one @[simp] @@ -1474,9 +1499,10 @@ theorem round_add_nat (x : α) (y : ℕ) : round (x + y) = round x + y := mod_cast round_add_int x y #align round_add_nat round_add_nat +-- See note [no_index around OfNat.ofNat] @[simp] theorem round_add_ofNat (x : α) (n : ℕ) [n.AtLeastTwo] : - round (x + OfNat.ofNat n) = round x + OfNat.ofNat n := + round (x + (no_index (OfNat.ofNat n))) = round x + OfNat.ofNat n := round_add_nat x n @[simp] @@ -1484,9 +1510,10 @@ theorem round_sub_nat (x : α) (y : ℕ) : round (x - y) = round x - y := mod_cast round_sub_int x y #align round_sub_nat round_sub_nat +-- See note [no_index around OfNat.ofNat] @[simp] theorem round_sub_ofNat (x : α) (n : ℕ) [n.AtLeastTwo] : - round (x - OfNat.ofNat n) = round x - OfNat.ofNat n := + round (x - (no_index (OfNat.ofNat n))) = round x - OfNat.ofNat n := round_sub_nat x n @[simp] @@ -1499,9 +1526,10 @@ theorem round_nat_add (x : α) (y : ℕ) : round ((y : α) + x) = y + round x := rw [add_comm, round_add_nat, add_comm] #align round_nat_add round_nat_add +-- See note [no_index around OfNat.ofNat] @[simp] theorem round_ofNat_add (n : ℕ) [n.AtLeastTwo] (x : α) : - round (OfNat.ofNat n + x) = OfNat.ofNat n + round x := + round ((no_index (OfNat.ofNat n)) + x) = OfNat.ofNat n + round x := round_nat_add x n theorem abs_sub_round_eq_min (x : α) : |x - round x| = min (fract x) (1 - fract x) := by diff --git a/Mathlib/Algebra/Order/Hom/Ring.lean b/Mathlib/Algebra/Order/Hom/Ring.lean index 3edbd20b00415..5ff8583808c1d 100644 --- a/Mathlib/Algebra/Order/Hom/Ring.lean +++ b/Mathlib/Algebra/Order/Hom/Ring.lean @@ -505,9 +505,8 @@ theorem symm_trans_self (e : α ≃+*o β) : e.symm.trans e = OrderRingIso.refl ext e.right_inv #align order_ring_iso.symm_trans_self OrderRingIso.symm_trans_self -theorem symm_bijective : Bijective (OrderRingIso.symm : α ≃+*o β → β ≃+*o α) := - ⟨fun f g h => f.symm_symm.symm.trans <| (congr_arg OrderRingIso.symm h).trans g.symm_symm, - fun f => ⟨f.symm, f.symm_symm⟩⟩ +theorem symm_bijective : Bijective (OrderRingIso.symm : (α ≃+*o β) → β ≃+*o α) := + Function.bijective_iff_has_inverse.mpr ⟨_, symm_symm, symm_symm⟩ #align order_ring_iso.symm_bijective OrderRingIso.symm_bijective end LE @@ -559,7 +558,7 @@ instance OrderRingHom.subsingleton [LinearOrderedField α] [LinearOrderedField Subsingleton (α →+*o β) := ⟨fun f g => by ext x - by_contra' h' : f x ≠ g x + by_contra! h' : f x ≠ g x wlog h : f x < g x generalizing α β with h₂ -- porting note: had to add the `generalizing` as there are random variables -- `F γ δ` flying around in context. diff --git a/Mathlib/Algebra/Order/LatticeGroup.lean b/Mathlib/Algebra/Order/LatticeGroup.lean index 81a384d8b8738..9b2e287d09b67 100644 --- a/Mathlib/Algebra/Order/LatticeGroup.lean +++ b/Mathlib/Algebra/Order/LatticeGroup.lean @@ -382,7 +382,7 @@ theorem one_le_abs [CovariantClass α α (· * ·) (· ≤ ·)] [CovariantClass α α (swap (· * ·)) (· ≤ ·)] (a : α) : 1 ≤ |a| := by apply pow_two_semiclosed _ _ - rw [abs_eq_sup_inv, pow_two, mul_sup, sup_mul, ←pow_two, mul_left_inv, sup_comm, ← sup_assoc] + rw [abs_eq_sup_inv, pow_two, mul_sup, sup_mul, ← pow_two, mul_left_inv, sup_comm, ← sup_assoc] apply le_sup_right -- The proof from Bourbaki A.VI.12 Prop 9 d) @@ -425,7 +425,7 @@ calc (a ⊔ b) / (a ⊓ b) = (a ⊔ b) * (a⁻¹ ⊔ b⁻¹) := by rw [div_eq_mul_inv, ← inv_inf_eq_sup_inv] _ = (a * a⁻¹ ⊔ b * a⁻¹) ⊔ (a * b⁻¹ ⊔ b * b⁻¹) := by rw [mul_sup, sup_mul, sup_mul] _ = (1 ⊔ b / a) ⊔ (a / b ⊔ 1) := by - rw [mul_right_inv, mul_right_inv, ←div_eq_mul_inv, ←div_eq_mul_inv] + rw [mul_right_inv, mul_right_inv, ← div_eq_mul_inv, ← div_eq_mul_inv] _ = 1 ⊔ b / a ⊔ (1 / (b / a) ⊔ 1) := by rw [one_div_div] _ = 1 ⊔ (b / a) ⊔ ((b / a)⁻¹ ⊔ 1) := by rw [inv_eq_one_div] _ = 1 ⊔ (((b / a) ⊔ (b / a)⁻¹) ⊔ 1) := by rw [sup_assoc, sup_assoc] @@ -657,12 +657,12 @@ variable [Semiring α] [Invertible (2 : α)] [Lattice β] [AddCommGroup β] [Mod lemma inf_eq_half_smul_add_sub_abs_sub (x y : β) : x ⊓ y = (⅟2 : α) • (x + y - |y - x|) := by - rw [←LatticeOrderedCommGroup.two_inf_eq_add_sub_abs_sub x y, two_smul, ←two_smul α, + rw [← LatticeOrderedCommGroup.two_inf_eq_add_sub_abs_sub x y, two_smul, ← two_smul α, smul_smul, invOf_mul_self, one_smul] lemma sup_eq_half_smul_add_add_abs_sub (x y : β) : x ⊔ y = (⅟2 : α) • (x + y + |y - x|) := by - rw [←LatticeOrderedCommGroup.two_sup_eq_add_add_abs_sub x y, two_smul, ←two_smul α, + rw [← LatticeOrderedCommGroup.two_sup_eq_add_add_abs_sub x y, two_smul, ← two_smul α, smul_smul, invOf_mul_self, one_smul] end invertible diff --git a/Mathlib/Algebra/Order/Module.lean b/Mathlib/Algebra/Order/Module.lean index 029187d36b8a7..ee95ae00050ea 100644 --- a/Mathlib/Algebra/Order/Module.lean +++ b/Mathlib/Algebra/Order/Module.lean @@ -71,7 +71,7 @@ theorem lt_of_smul_lt_smul_of_nonpos (h : c • a < c • b) (hc : c ≤ 0) : b #align lt_of_smul_lt_smul_of_nonpos lt_of_smul_lt_smul_of_nonpos lemma smul_le_smul_of_nonneg_right (h : c ≤ d) (hb : 0 ≤ b) : c • b ≤ d • b := by - rw [←sub_nonneg, ←sub_smul]; exact smul_nonneg (sub_nonneg.2 h) hb + rw [← sub_nonneg, ← sub_smul]; exact smul_nonneg (sub_nonneg.2 h) hb lemma smul_le_smul (hcd : c ≤ d) (hab : a ≤ b) (ha : 0 ≤ a) (hd : 0 ≤ d) : c • a ≤ d • b := (smul_le_smul_of_nonneg_right hcd ha).trans $ smul_le_smul_of_nonneg_left hab hd @@ -248,23 +248,23 @@ lemma smul_nonneg_iff : 0 ≤ a • b ↔ 0 ≤ a ∧ 0 ≤ b ∨ a ≤ 0 ∧ b fun h ↦ h.elim (and_imp.2 smul_nonneg) (and_imp.2 smul_nonneg_of_nonpos_of_nonpos)⟩ lemma smul_nonpos_iff : a • b ≤ 0 ↔ 0 ≤ a ∧ b ≤ 0 ∨ a ≤ 0 ∧ 0 ≤ b := by - rw [←neg_nonneg, ←smul_neg, smul_nonneg_iff, neg_nonneg, neg_nonpos] + rw [← neg_nonneg, ← smul_neg, smul_nonneg_iff, neg_nonneg, neg_nonpos] lemma smul_nonneg_iff_pos_imp_nonneg : 0 ≤ a • b ↔ (0 < a → 0 ≤ b) ∧ (0 < b → 0 ≤ a) := by refine smul_nonneg_iff.trans ?_ - simp_rw [←not_le, ←or_iff_not_imp_left] + simp_rw [← not_le, ← or_iff_not_imp_left] have := le_total a 0 have := le_total b 0 tauto lemma smul_nonneg_iff_neg_imp_nonpos : 0 ≤ a • b ↔ (a < 0 → b ≤ 0) ∧ (b < 0 → a ≤ 0) := by - rw [←neg_smul_neg, smul_nonneg_iff_pos_imp_nonneg]; simp only [neg_pos, neg_nonneg] + rw [← neg_smul_neg, smul_nonneg_iff_pos_imp_nonneg]; simp only [neg_pos, neg_nonneg] lemma smul_nonpos_iff_pos_imp_nonpos : a • b ≤ 0 ↔ (0 < a → b ≤ 0) ∧ (b < 0 → 0 ≤ a) := by - rw [←neg_nonneg, ←smul_neg, smul_nonneg_iff_pos_imp_nonneg]; simp only [neg_pos, neg_nonneg] + rw [← neg_nonneg, ← smul_neg, smul_nonneg_iff_pos_imp_nonneg]; simp only [neg_pos, neg_nonneg] lemma smul_nonpos_iff_neg_imp_nonneg : a • b ≤ 0 ↔ (a < 0 → 0 ≤ b) ∧ (0 < b → a ≤ 0) := by - rw [←neg_nonneg, ←neg_smul, smul_nonneg_iff_pos_imp_nonneg]; simp only [neg_pos, neg_nonneg] + rw [← neg_nonneg, ← neg_smul, smul_nonneg_iff_pos_imp_nonneg]; simp only [neg_pos, neg_nonneg] end LinearOrderedRing diff --git a/Mathlib/Algebra/Order/Monoid/Canonical/Defs.lean b/Mathlib/Algebra/Order/Monoid/Canonical/Defs.lean index 81d5b488e4a6f..1985000eae8b8 100644 --- a/Mathlib/Algebra/Order/Monoid/Canonical/Defs.lean +++ b/Mathlib/Algebra/Order/Monoid/Canonical/Defs.lean @@ -275,7 +275,7 @@ theorem le_mul_right (h : a ≤ b) : a ≤ b * c := @[to_additive] theorem lt_iff_exists_mul [CovariantClass α α (· * ·) (· < ·)] : a < b ↔ ∃ c > 1, b = a * c := by - rw [lt_iff_le_and_ne, le_iff_exists_mul, ←exists_and_right] + rw [lt_iff_le_and_ne, le_iff_exists_mul, ← exists_and_right] apply exists_congr intro c rw [and_comm, and_congr_left_iff, gt_iff_lt] diff --git a/Mathlib/Algebra/Order/Monoid/WithTop.lean b/Mathlib/Algebra/Order/Monoid/WithTop.lean index 9fce95d0880fa..befe9c4750044 100644 --- a/Mathlib/Algebra/Order/Monoid/WithTop.lean +++ b/Mathlib/Algebra/Order/Monoid/WithTop.lean @@ -185,7 +185,7 @@ theorem add_right_cancel_iff [IsRightCancelAdd α] (ha : a ≠ ⊤) : b + a = c obtain rfl | hb := (eq_or_ne b ⊤) · rw [top_add, eq_comm, WithTop.add_coe_eq_top_iff, eq_comm] lift b to α using hb - simp_rw [←WithTop.coe_add, eq_comm, WithTop.add_eq_coe, coe_eq_coe, exists_and_left, + simp_rw [← WithTop.coe_add, eq_comm, WithTop.add_eq_coe, coe_eq_coe, exists_and_left, exists_eq_left, add_left_inj, exists_eq_right, eq_comm] theorem add_right_cancel [IsRightCancelAdd α] (ha : a ≠ ⊤) (h : b + a = c + a) : b = c := @@ -196,7 +196,7 @@ theorem add_left_cancel_iff [IsLeftCancelAdd α] (ha : a ≠ ⊤) : a + b = a + obtain rfl | hb := (eq_or_ne b ⊤) · rw [add_top, eq_comm, WithTop.coe_add_eq_top_iff, eq_comm] lift b to α using hb - simp_rw [←WithTop.coe_add, eq_comm, WithTop.add_eq_coe, eq_comm, coe_eq_coe, + simp_rw [← WithTop.coe_add, eq_comm, WithTop.add_eq_coe, eq_comm, coe_eq_coe, exists_and_left, exists_eq_left', add_right_inj, exists_eq_right'] theorem add_left_cancel [IsLeftCancelAdd α] (ha : a ≠ ⊤) (h : a + b = a + c) : b = c := diff --git a/Mathlib/Algebra/Order/Monovary.lean b/Mathlib/Algebra/Order/Monovary.lean index 31ea4fb4f84c0..30fe1693595b8 100644 --- a/Mathlib/Algebra/Order/Monovary.lean +++ b/Mathlib/Algebra/Order/Monovary.lean @@ -346,7 +346,7 @@ lemma monovaryOn_iff_smul_rearrangement : MonovaryOn f g s ↔ ∀ ⦃i⦄, i ∈ s → ∀ ⦃j⦄, j ∈ s → f i • g j + f j • g i ≤ f i • g i + f j • g j := monovaryOn_iff_forall_smul_nonneg.trans $ forall₄_congr $ fun i _ j _ ↦ by - simp [smul_sub, sub_smul, ←add_sub_right_comm, le_sub_iff_add_le, add_comm (f i • g i), + simp [smul_sub, sub_smul, ← add_sub_right_comm, le_sub_iff_add_le, add_comm (f i • g i), add_comm (f i • g j)] /-- Two functions antivary iff the rearrangement inequality holds. -/ diff --git a/Mathlib/Algebra/Order/Ring/Lemmas.lean b/Mathlib/Algebra/Order/Ring/Lemmas.lean index 7ed5560aa06fc..b3845ffc35bd0 100644 --- a/Mathlib/Algebra/Order/Ring/Lemmas.lean +++ b/Mathlib/Algebra/Order/Ring/Lemmas.lean @@ -352,7 +352,7 @@ theorem mul_neg_of_pos_of_neg [PosMulStrictMono α] (ha : 0 < a) (hb : b < 0) : @[simp] theorem zero_lt_mul_left [PosMulStrictMono α] [PosMulReflectLT α] (h : 0 < c) : 0 < c * b ↔ 0 < b := by - rw [←mul_zero c, mul_lt_mul_left h] + rw [← mul_zero c, mul_lt_mul_left h] simp #align zero_lt_mul_left zero_lt_mul_left @@ -368,7 +368,7 @@ theorem mul_neg_of_neg_of_pos [MulPosStrictMono α] (ha : a < 0) (hb : 0 < b) : @[simp] theorem zero_lt_mul_right [MulPosStrictMono α] [MulPosReflectLT α] (h : 0 < c) : 0 < b * c ↔ 0 < b := by - rw [←zero_mul c, mul_lt_mul_right h] + rw [← zero_mul c, mul_lt_mul_right h] simp #align zero_lt_mul_right zero_lt_mul_right @@ -452,7 +452,7 @@ theorem posMulMono_iff_covariant_pos : ⟨@PosMulMono.to_covariantClass_pos_mul_le _ _ _ _, fun h => ⟨fun a b c h => by obtain ha | ha := a.prop.eq_or_lt - · simp [←ha] + · simp [← ha] · exact @CovariantClass.elim α>0 α (fun x y => x * y) (· ≤ ·) _ ⟨_, ha⟩ _ _ h ⟩⟩ #align pos_mul_mono_iff_covariant_pos posMulMono_iff_covariant_pos @@ -461,7 +461,7 @@ theorem mulPosMono_iff_covariant_pos : ⟨@MulPosMono.to_covariantClass_pos_mul_le _ _ _ _, fun h => ⟨fun a b c h => by obtain ha | ha := a.prop.eq_or_lt - · simp [←ha] + · simp [← ha] · exact @CovariantClass.elim α>0 α (fun x y => y * x) (· ≤ ·) _ ⟨_, ha⟩ _ _ h ⟩⟩ #align mul_pos_mono_iff_covariant_pos mulPosMono_iff_covariant_pos @@ -470,7 +470,7 @@ theorem posMulReflectLT_iff_contravariant_pos : ⟨@PosMulReflectLT.to_contravariantClass_pos_mul_lt _ _ _ _, fun h => ⟨fun a b c h => by obtain ha | ha := a.prop.eq_or_lt - · simp [←ha] at h + · simp [← ha] at h · exact @ContravariantClass.elim α>0 α (fun x y => x * y) (· < ·) _ ⟨_, ha⟩ _ _ h ⟩⟩ #align pos_mul_reflect_lt_iff_contravariant_pos posMulReflectLT_iff_contravariant_pos @@ -479,7 +479,7 @@ theorem mulPosReflectLT_iff_contravariant_pos : ⟨@MulPosReflectLT.to_contravariantClass_pos_mul_lt _ _ _ _, fun h => ⟨fun a b c h => by obtain ha | ha := a.prop.eq_or_lt - · simp [←ha] at h + · simp [← ha] at h · exact @ContravariantClass.elim α>0 α (fun x y => y * x) (· < ·) _ ⟨_, ha⟩ _ _ h ⟩⟩ #align mul_pos_reflect_lt_iff_contravariant_pos mulPosReflectLT_iff_contravariant_pos diff --git a/Mathlib/Algebra/Order/Ring/Star.lean b/Mathlib/Algebra/Order/Ring/Star.lean index 5294a90b0d9b1..13b4b35f6981f 100644 --- a/Mathlib/Algebra/Order/Ring/Star.lean +++ b/Mathlib/Algebra/Order/Ring/Star.lean @@ -34,7 +34,7 @@ example {R : Type*} [OrderedSemiring R] [StarOrderedRing R] {x y : R} (hx : 0 induction hz using AddSubmonoid.closure_induction' with | Hs _ h => obtain ⟨x, rfl⟩ := h; simp | H1 => simp - | Hmul x hx y hy => simp only [←nonneg_iff] at hx hy; aesop + | Hmul x hx y hy => simp only [← nonneg_iff] at hx hy; aesop -- `0 ≤ y * x`, and hence `y * x` is self-adjoint have := this _ <| mul_nonneg hy hx aesop @@ -47,11 +47,11 @@ private lemma mul_le_mul_of_nonneg_left {R : Type*} [CommSemiring R] [PartialOrd induction hc using AddSubmonoid.closure_induction' with | Hs _ h => obtain ⟨x, rfl⟩ := h - simp_rw [mul_assoc, mul_comm x, ←mul_assoc] + simp_rw [mul_assoc, mul_comm x, ← mul_assoc] exact conjugate_le_conjugate hab x | H1 => simp | Hmul x hx y hy => - simp only [←nonneg_iff, add_mul] at hx hy ⊢ + simp only [← nonneg_iff, add_mul] at hx hy ⊢ apply add_le_add <;> aesop /-- A commutative star-ordered semiring is an ordered semiring. diff --git a/Mathlib/Algebra/Order/Sub/Defs.lean b/Mathlib/Algebra/Order/Sub/Defs.lean index 165415857bb95..964d88879b82a 100644 --- a/Mathlib/Algebra/Order/Sub/Defs.lean +++ b/Mathlib/Algebra/Order/Sub/Defs.lean @@ -274,7 +274,7 @@ theorem tsub_add_eq_tsub_tsub_swap (a b c : α) : a - (b + c) = a - c - b := by #align tsub_add_eq_tsub_tsub_swap tsub_add_eq_tsub_tsub_swap theorem tsub_right_comm : a - b - c = a - c - b := by - rw [←tsub_add_eq_tsub_tsub, tsub_add_eq_tsub_tsub_swap] + rw [← tsub_add_eq_tsub_tsub, tsub_add_eq_tsub_tsub_swap] #align tsub_right_comm tsub_right_comm /-! ### Lemmas that assume that an element is `AddLECancellable`. -/ diff --git a/Mathlib/Algebra/Order/ToIntervalMod.lean b/Mathlib/Algebra/Order/ToIntervalMod.lean index ab267a52feae6..d568d1ebef8d4 100644 --- a/Mathlib/Algebra/Order/ToIntervalMod.lean +++ b/Mathlib/Algebra/Order/ToIntervalMod.lean @@ -877,7 +877,7 @@ private theorem toIxxMod_cyclic_left {x₁ x₂ x₃ : α} (h : toIcoMod hp x₁ private theorem toIxxMod_antisymm (h₁₂₃ : toIcoMod hp a b ≤ toIocMod hp a c) (h₁₃₂ : toIcoMod hp a c ≤ toIocMod hp a b) : b ≡ a [PMOD p] ∨ c ≡ b [PMOD p] ∨ a ≡ c [PMOD p] := by - by_contra' h + by_contra! h rw [modEq_comm] at h rw [← (not_modEq_iff_toIcoMod_eq_toIocMod hp).mp h.2.2] at h₁₂₃ rw [← (not_modEq_iff_toIcoMod_eq_toIocMod hp).mp h.1] at h₁₃₂ diff --git a/Mathlib/Algebra/Order/WithZero.lean b/Mathlib/Algebra/Order/WithZero.lean index a9be045b370ee..74102a839763e 100644 --- a/Mathlib/Algebra/Order/WithZero.lean +++ b/Mathlib/Algebra/Order/WithZero.lean @@ -55,7 +55,7 @@ instance instLinearOrderedCommMonoidWithZeroMultiplicativeOrderDual instance [LinearOrderedAddCommGroupWithTop α] : LinearOrderedCommGroupWithZero (Multiplicative αᵒᵈ) := { Multiplicative.divInvMonoid, instLinearOrderedCommMonoidWithZeroMultiplicativeOrderDual, - instNontrivialMultiplicative with + Multiplicative.instNontrivial with inv_zero := @LinearOrderedAddCommGroupWithTop.neg_top _ (_) mul_inv_cancel := @LinearOrderedAddCommGroupWithTop.add_neg_cancel _ (_) } @@ -283,6 +283,6 @@ theorem OrderIso.mulRight₀'_symm {a : α} (ha : a ≠ 0) : instance : LinearOrderedAddCommGroupWithTop (Additive αᵒᵈ) := { Additive.subNegMonoid, instLinearOrderedAddCommMonoidWithTopAdditiveOrderDual, - instNontrivialAdditive with + Additive.instNontrivial with neg_top := @inv_zero _ (_) add_neg_cancel := fun a ha ↦ mul_inv_cancel (id ha : Additive.toMul a ≠ 0) } diff --git a/Mathlib/Algebra/Quandle.lean b/Mathlib/Algebra/Quandle.lean index 4143d6ddcc789..72522e50dd61b 100644 --- a/Mathlib/Algebra/Quandle.lean +++ b/Mathlib/Algebra/Quandle.lean @@ -160,15 +160,15 @@ identity holds. -/ lemma act_act_self_eq (x y : S) : (x ◃ y) ◃ x = x ◃ y := by have h : (x ◃ y) ◃ x = (x ◃ y) ◃ (x ◃ 1) := by rw [act_one] - rw [h, ←Shelf.self_distrib, act_one] + rw [h, ← Shelf.self_distrib, act_one] #align unital_shelf.act_act_self_eq UnitalShelf.act_act_self_eq -lemma act_idem (x : S) : (x ◃ x) = x := by rw [←act_one x, ←Shelf.self_distrib, act_one, act_one] +lemma act_idem (x : S) : (x ◃ x) = x := by rw [← act_one x, ← Shelf.self_distrib, act_one, act_one] #align unital_shelf.act_idem UnitalShelf.act_idem lemma act_self_act_eq (x y : S) : x ◃ (x ◃ y) = x ◃ y := by have h : x ◃ (x ◃ y) = (x ◃ 1) ◃ (x ◃ y) := by rw [act_one] - rw [h, ←Shelf.self_distrib, one_act] + rw [h, ← Shelf.self_distrib, one_act] #align unital_shelf.act_self_act_eq UnitalShelf.act_self_act_eq /-- diff --git a/Mathlib/Algebra/Quaternion.lean b/Mathlib/Algebra/Quaternion.lean index 6fa99a365571a..e6978519610c4 100644 --- a/Mathlib/Algebra/Quaternion.lean +++ b/Mathlib/Algebra/Quaternion.lean @@ -1413,10 +1413,10 @@ instance : DivisionRing ℍ[R] := { Quaternion.instGroupWithZero, Quaternion.instRing with ratCast_mk := fun n d hd h => by - rw [←coe_rat_cast, Rat.cast_mk', coe_mul, coe_int_cast, coe_inv, coe_nat_cast] + rw [← coe_rat_cast, Rat.cast_mk', coe_mul, coe_int_cast, coe_inv, coe_nat_cast] qsmul := (· • ·) qsmul_eq_mul' := fun q x => by - rw [←coe_rat_cast, coe_mul_eq_smul] + rw [← coe_rat_cast, coe_mul_eq_smul] ext <;> exact DivisionRing.qsmul_eq_mul' _ _ } --@[simp] Porting note: `simp` can prove it diff --git a/Mathlib/Algebra/Regular/SMul.lean b/Mathlib/Algebra/Regular/SMul.lean index 624f4175b549b..938a3342056db 100644 --- a/Mathlib/Algebra/Regular/SMul.lean +++ b/Mathlib/Algebra/Regular/SMul.lean @@ -74,7 +74,7 @@ element, then `b` is `M`-regular. -/ theorem of_smul (a : R) (ab : IsSMulRegular M (a • s)) : IsSMulRegular M s := @Function.Injective.of_comp _ _ _ (fun m : M => a • m) _ fun c d cd => by dsimp only [Function.comp_def] at cd - rw [←smul_assoc, ←smul_assoc] at cd + rw [← smul_assoc, ← smul_assoc] at cd exact ab cd #align is_smul_regular.of_smul IsSMulRegular.of_smul diff --git a/Mathlib/Algebra/Ring/Basic.lean b/Mathlib/Algebra/Ring/Basic.lean index 8db18855d12e3..9603061496fb3 100644 --- a/Mathlib/Algebra/Ring/Basic.lean +++ b/Mathlib/Algebra/Ring/Basic.lean @@ -147,7 +147,7 @@ theorem succ_ne_self [NonAssocRing α] [Nontrivial α] (a : α) : a + 1 ≠ a := #align succ_ne_self succ_ne_self theorem pred_ne_self [NonAssocRing α] [Nontrivial α] (a : α) : a - 1 ≠ a := fun h ↦ - one_ne_zero (neg_injective ((add_right_inj a).mp (by simp [←sub_eq_add_neg, h]))) + one_ne_zero (neg_injective ((add_right_inj a).mp (by simp [← sub_eq_add_neg, h]))) #align pred_ne_self pred_ne_self section NoZeroDivisors diff --git a/Mathlib/Algebra/Ring/Defs.lean b/Mathlib/Algebra/Ring/Defs.lean index bd5e50aeb6e2a..d6abe027ec83e 100644 --- a/Mathlib/Algebra/Ring/Defs.lean +++ b/Mathlib/Algebra/Ring/Defs.lean @@ -236,6 +236,9 @@ theorem ite_and_mul_zero {α : Type*} [MulZeroClass α] (P Q : Prop) [Decidable simp only [← ite_and, ite_mul, mul_ite, mul_zero, zero_mul, and_comm] #align ite_and_mul_zero ite_and_mul_zero +/-- A not-necessarily-unital, not-necessarily-associative, but commutative semiring. -/ +class NonUnitalNonAssocCommSemiring (α : Type u) extends NonUnitalNonAssocSemiring α, CommMagma α + /-- A non-unital commutative semiring is a `NonUnitalSemiring` with commutative multiplication. In other words, it is a type with the following structures: additive commutative monoid (`AddCommMonoid`), commutative semigroup (`CommSemigroup`), distributive laws (`Distrib`), and @@ -421,8 +424,13 @@ instance (priority := 200) : Semiring α := end Ring +/-- A non-unital non-associative commutative ring is a `NonUnitalNonAssocRing` with commutative +multiplication. -/ +class NonUnitalNonAssocCommRing (α : Type u) + extends NonUnitalNonAssocRing α, NonUnitalNonAssocCommSemiring α + /-- A non-unital commutative ring is a `NonUnitalRing` with commutative multiplication. -/ -class NonUnitalCommRing (α : Type u) extends NonUnitalRing α, CommSemigroup α +class NonUnitalCommRing (α : Type u) extends NonUnitalRing α, NonUnitalNonAssocCommRing α #align non_unital_comm_ring NonUnitalCommRing -- see Note [lower instance priority] diff --git a/Mathlib/Algebra/Ring/Equiv.lean b/Mathlib/Algebra/Ring/Equiv.lean index bc348ff97752d..6b4fc70930171 100644 --- a/Mathlib/Algebra/Ring/Equiv.lean +++ b/Mathlib/Algebra/Ring/Equiv.lean @@ -295,8 +295,8 @@ theorem coe_toEquiv_symm (e : R ≃+* S) : (e.symm : S ≃ R) = (e : R ≃ S).sy rfl #align ring_equiv.coe_to_equiv_symm RingEquiv.coe_toEquiv_symm -theorem symm_bijective : Function.Bijective (RingEquiv.symm : R ≃+* S → S ≃+* R) := - Equiv.bijective ⟨RingEquiv.symm, RingEquiv.symm, symm_symm, symm_symm⟩ +theorem symm_bijective : Function.Bijective (RingEquiv.symm : (R ≃+* S) → S ≃+* R) := + Function.bijective_iff_has_inverse.mpr ⟨_, symm_symm, symm_symm⟩ #align ring_equiv.symm_bijective RingEquiv.symm_bijective @[simp] diff --git a/Mathlib/Algebra/Ring/Units.lean b/Mathlib/Algebra/Ring/Units.lean index 43720b18af397..c1e778d912710 100644 --- a/Mathlib/Algebra/Ring/Units.lean +++ b/Mathlib/Algebra/Ring/Units.lean @@ -114,7 +114,7 @@ theorem divp_add_divp [CommRing α] (a b : α) (u₁ u₂ : αˣ) : a /ₚ u₁ + b /ₚ u₂ = (a * u₂ + u₁ * b) /ₚ (u₁ * u₂) := by simp only [divp, add_mul, mul_inv_rev, val_mul] rw [mul_comm (↑u₁ * b), mul_comm b] - rw [←mul_assoc, ←mul_assoc, mul_assoc a, mul_assoc (↑u₂⁻¹ : α), mul_inv, inv_mul, mul_one, + rw [← mul_assoc, ← mul_assoc, mul_assoc a, mul_assoc (↑u₂⁻¹ : α), mul_inv, inv_mul, mul_one, mul_one] -- porting note: `assoc_rw` not ported: `assoc_rw [mul_inv, mul_inv, mul_one, mul_one]` #align units.divp_add_divp Units.divp_add_divp diff --git a/Mathlib/Algebra/SMulWithZero.lean b/Mathlib/Algebra/SMulWithZero.lean index 89befdee8ae8f..1e79f2f3e9a2a 100644 --- a/Mathlib/Algebra/SMulWithZero.lean +++ b/Mathlib/Algebra/SMulWithZero.lean @@ -105,7 +105,7 @@ protected def Function.Surjective.smulWithZero (f : ZeroHom M M') (hf : Function zero_smul m := by rcases hf m with ⟨x, rfl⟩ simp [← smul] - smul_zero c := by rw [←f.map_zero, ←smul, smul_zero] + smul_zero c := by rw [← f.map_zero, ← smul, smul_zero] #align function.surjective.smul_with_zero Function.Surjective.smulWithZero variable (M) @@ -165,7 +165,8 @@ instance MonoidWithZero.toOppositeMulActionWithZero : MulActionWithZero Rᵐᵒ protected lemma MulActionWithZero.subsingleton [MulActionWithZero R M] [Subsingleton R] : Subsingleton M := - ⟨λ x y => by rw [←one_smul R x, ←one_smul R y, Subsingleton.elim (1 : R) 0, zero_smul, zero_smul]⟩ + ⟨λ x y => by + rw [← one_smul R x, ← one_smul R y, Subsingleton.elim (1 : R) 0, zero_smul, zero_smul]⟩ #align mul_action_with_zero.subsingleton MulActionWithZero.subsingleton protected lemma MulActionWithZero.nontrivial diff --git a/Mathlib/Algebra/Squarefree.lean b/Mathlib/Algebra/Squarefree.lean index 2865eb8152059..e9fb2fe7d16a4 100644 --- a/Mathlib/Algebra/Squarefree.lean +++ b/Mathlib/Algebra/Squarefree.lean @@ -109,7 +109,7 @@ theorem squarefree_iff_multiplicity_le_one (r : R) : Squarefree r ↔ ∀ x : R, multiplicity x r ≤ 1 ∨ IsUnit x := by refine' forall_congr' fun a => _ rw [← sq, pow_dvd_iff_le_multiplicity, or_iff_not_imp_left, not_le, imp_congr _ Iff.rfl] - rw [←one_add_one_eq_two] + rw [← one_add_one_eq_two] simpa using PartENat.add_one_le_iff_lt (PartENat.natCast_ne_top 1) #align multiplicity.squarefree_iff_multiplicity_le_one multiplicity.squarefree_iff_multiplicity_le_one diff --git a/Mathlib/Algebra/Star/NonUnitalSubalgebra.lean b/Mathlib/Algebra/Star/NonUnitalSubalgebra.lean index f280fb732582a..909408e907a6f 100644 --- a/Mathlib/Algebra/Star/NonUnitalSubalgebra.lean +++ b/Mathlib/Algebra/Star/NonUnitalSubalgebra.lean @@ -694,11 +694,11 @@ theorem toNonUnitalSubalgebra_eq_top {S : NonUnitalStarSubalgebra R A} : NonUnitalStarSubalgebra.toNonUnitalSubalgebra_injective.eq_iff' top_toNonUnitalSubalgebra theorem mem_sup_left {S T : NonUnitalStarSubalgebra R A} : ∀ {x : A}, x ∈ S → x ∈ S ⊔ T := by - rw [←SetLike.le_def] + rw [← SetLike.le_def] exact le_sup_left theorem mem_sup_right {S T : NonUnitalStarSubalgebra R A} : ∀ {x : A}, x ∈ T → x ∈ S ⊔ T := by - rw [←SetLike.le_def] + rw [← SetLike.le_def] exact le_sup_right theorem mul_mem_sup {S T : NonUnitalStarSubalgebra R A} {x y : A} (hx : x ∈ S) (hy : y ∈ T) : diff --git a/Mathlib/Algebra/Star/StarAlgHom.lean b/Mathlib/Algebra/Star/StarAlgHom.lean index 3732d611a92d6..8a9b7b3362da8 100644 --- a/Mathlib/Algebra/Star/StarAlgHom.lean +++ b/Mathlib/Algebra/Star/StarAlgHom.lean @@ -867,7 +867,7 @@ theorem symm_symm (e : A ≃⋆ₐ[R] B) : e.symm.symm = e := by #align star_alg_equiv.symm_symm StarAlgEquiv.symm_symm theorem symm_bijective : Function.Bijective (symm : (A ≃⋆ₐ[R] B) → B ≃⋆ₐ[R] A) := - Equiv.bijective ⟨symm, symm, symm_symm, symm_symm⟩ + Function.bijective_iff_has_inverse.mpr ⟨_, symm_symm, symm_symm⟩ #align star_alg_equiv.symm_bijective StarAlgEquiv.symm_bijective -- porting note: doesn't align with Mathlib 3 because `StarAlgEquiv.mk` has a new signature diff --git a/Mathlib/Algebra/Star/Subalgebra.lean b/Mathlib/Algebra/Star/Subalgebra.lean index f5f51ebc32ac1..85a26c0f1e140 100644 --- a/Mathlib/Algebra/Star/Subalgebra.lean +++ b/Mathlib/Algebra/Star/Subalgebra.lean @@ -736,7 +736,7 @@ variable [hF : StarAlgHomClass F R A B] (f g : F) /-- The equalizer of two star `R`-algebra homomorphisms. -/ def equalizer : StarSubalgebra R A := { toSubalgebra := AlgHom.equalizer (f : A →ₐ[R] B) g - star_mem' := @fun a (ha : f a = g a) => by simpa only [←map_star] using congrArg star ha } + star_mem' := @fun a (ha : f a = g a) => by simpa only [← map_star] using congrArg star ha } -- porting note: much like `StarSubalgebra.copy` the old proof was broken and hard to fix #align star_alg_hom.equalizer StarAlgHom.equalizer diff --git a/Mathlib/Algebra/Symmetrized.lean b/Mathlib/Algebra/Symmetrized.lean index 5754195404da2..8760258d3006c 100644 --- a/Mathlib/Algebra/Symmetrized.lean +++ b/Mathlib/Algebra/Symmetrized.lean @@ -336,8 +336,10 @@ theorem mul_comm [Mul α] [AddCommSemigroup α] [One α] [OfNat α 2] [Invertibl a * b = b * a := by rw [mul_def, mul_def, add_comm] #align sym_alg.mul_comm SymAlg.mul_comm -instance [Ring α] [Invertible (2 : α)] : IsCommJordan αˢʸᵐ where +instance [Ring α] [Invertible (2 : α)] : CommMagma αˢʸᵐ where mul_comm := SymAlg.mul_comm + +instance [Ring α] [Invertible (2 : α)] : IsCommJordan αˢʸᵐ where lmul_comm_rmul_rmul a b := by have commute_half_left := fun a : α => by -- Porting note: mathlib3 used `bit0_left` diff --git a/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean b/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean index 6f3e24b709846..5c0c54363c4e1 100644 --- a/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean +++ b/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean @@ -197,7 +197,7 @@ theorem toΓSpecSheafedSpace_app_eq : ((TopCat.Sheaf.pushforward _ X.toΓSpecBase).obj X.𝒪) isBasis_basic_opens X.toΓSpecCBasicOpens r dsimp at this - rw [←this] + rw [← this] dsimp #align algebraic_geometry.LocallyRingedSpace.to_Γ_Spec_SheafedSpace_app_eq AlgebraicGeometry.LocallyRingedSpace.toΓSpecSheafedSpace_app_eq @@ -219,7 +219,7 @@ theorem toStalk_stalkMap_toΓSpec (x : X) : ⟨X.toΓSpecFun x, by rw [basicOpen_one]; trivial⟩] rw [← Category.assoc, Category.assoc (toOpen _ _)] erw [stalkFunctor_map_germ] - -- Porting note : was `rw [←assoc, toΓSpecSheafedSpace_app_spec]`, but Lean did not like it. + -- Porting note : was `rw [← assoc, toΓSpecSheafedSpace_app_spec]`, but Lean did not like it. rw [toΓSpecSheafedSpace_app_spec_assoc] unfold ΓToStalk rw [← stalkPushforward_germ _ X.toΓSpecBase X.presheaf ⊤] diff --git a/Mathlib/AlgebraicGeometry/Morphisms/QuasiSeparated.lean b/Mathlib/AlgebraicGeometry/Morphisms/QuasiSeparated.lean index 3b87d06632ca8..9b251d3054082 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/QuasiSeparated.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/QuasiSeparated.lean @@ -291,7 +291,7 @@ theorem quasiSeparatedOfComp {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) [H : Q rintro ⟨i, j⟩; dsimp at i j -- replace H := H (Scheme.OpenCover.pullbackCover (Scheme.affineCover Z) g) i specialize H _ i - -- rw [←isQuasiSeparated_iff_quasiSeparatedSpace] at H + -- rw [← isQuasiSeparated_iff_quasiSeparatedSpace] at H refine @quasiSeparatedSpace_of_quasiSeparated _ _ ?_ H ?_ · exact pullback.map _ _ _ _ (𝟙 _) _ _ (by simp) (Category.comp_id _) ≫ (pullbackRightPullbackFstIso g (Z.affineCover.map i) f).hom diff --git a/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean b/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean index 6819c19d1a465..6c7f37b33559d 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean @@ -134,7 +134,7 @@ theorem StableUnderBaseChange.Γ_pullback_fst (hP : StableUnderBaseChange @P) (h simp only [Quiver.Hom.unop_op, Functor.comp_map, AffineScheme.forgetToScheme_map, Functor.op_map] at this rw [← this, hP'.cancel_right_isIso, - ←pushoutIsoUnopPullback_inl_hom (Quiver.Hom.unop _) (Quiver.Hom.unop _), + ← pushoutIsoUnopPullback_inl_hom (Quiver.Hom.unop _) (Quiver.Hom.unop _), hP'.cancel_right_isIso] exact hP.pushout_inl _ hP' _ _ H #align ring_hom.stable_under_base_change.Γ_pullback_fst RingHom.StableUnderBaseChange.Γ_pullback_fst diff --git a/Mathlib/AlgebraicGeometry/OpenImmersion.lean b/Mathlib/AlgebraicGeometry/OpenImmersion.lean index a933b9e7736c7..82829800fab70 100644 --- a/Mathlib/AlgebraicGeometry/OpenImmersion.lean +++ b/Mathlib/AlgebraicGeometry/OpenImmersion.lean @@ -267,7 +267,7 @@ theorem affineBasisCover_is_basis (X : Scheme) : TopologicalSpace.IsTopologicalBasis {x : Set X | ∃ a : X.affineBasisCover.J, x = Set.range (X.affineBasisCover.map a).1.base} := by - apply TopologicalSpace.isTopologicalBasis_of_open_of_nhds + apply TopologicalSpace.isTopologicalBasis_of_isOpen_of_nhds · rintro _ ⟨a, rfl⟩ exact IsOpenImmersion.open_range (X.affineBasisCover.map a) · rintro a U haU hU diff --git a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean index 3b932ca68adfc..b31c49ee9770f 100644 --- a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean +++ b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean @@ -808,7 +808,7 @@ theorem basicOpen_pow (f : R) (n : ℕ) (hn : 0 < n) : basicOpen (f ^ n) = basic theorem isTopologicalBasis_basic_opens : TopologicalSpace.IsTopologicalBasis (Set.range fun r : R => (basicOpen r : Set (PrimeSpectrum R))) := by - apply TopologicalSpace.isTopologicalBasis_of_open_of_nhds + apply TopologicalSpace.isTopologicalBasis_of_isOpen_of_nhds · rintro _ ⟨r, rfl⟩ exact isOpen_basicOpen · rintro p U hp ⟨s, hs⟩ diff --git a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Topology.lean b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Topology.lean index e1761f387f8c0..0f073ade8d882 100644 --- a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Topology.lean +++ b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Topology.lean @@ -453,7 +453,7 @@ theorem basicOpen_eq_union_of_projection (f : A) : theorem isTopologicalBasis_basic_opens : TopologicalSpace.IsTopologicalBasis (Set.range fun r : A => (basicOpen 𝒜 r : Set (ProjectiveSpectrum 𝒜))) := by - apply TopologicalSpace.isTopologicalBasis_of_open_of_nhds + apply TopologicalSpace.isTopologicalBasis_of_isOpen_of_nhds · rintro _ ⟨r, rfl⟩ exact isOpen_basicOpen 𝒜 · rintro p U hp ⟨s, hs⟩ diff --git a/Mathlib/AlgebraicGeometry/Properties.lean b/Mathlib/AlgebraicGeometry/Properties.lean index ddaaa15ca1518..cc8941ce209f7 100644 --- a/Mathlib/AlgebraicGeometry/Properties.lean +++ b/Mathlib/AlgebraicGeometry/Properties.lean @@ -261,7 +261,7 @@ theorem isIntegralOfIsIrreducibleIsReduced [IsReduced X] [H : IrreducibleSpace X (X.toLocallyRingedSpace.toSheafedSpace.toPresheafedSpace.presheaf.obj (op U)) := by refine' ⟨fun {a b} e => _⟩ simp_rw [← basicOpen_eq_bot_iff, ← Opens.not_nonempty_iff_eq_bot] - by_contra' h + by_contra! h obtain ⟨_, ⟨x, hx₁, rfl⟩, ⟨x, hx₂, e'⟩⟩ := nonempty_preirreducible_inter (X.basicOpen a).2 (X.basicOpen b).2 h.1 h.2 replace e' := Subtype.eq e' diff --git a/Mathlib/AlgebraicGeometry/Spec.lean b/Mathlib/AlgebraicGeometry/Spec.lean index 7bc32cebef6a3..4e103f98edae2 100644 --- a/Mathlib/AlgebraicGeometry/Spec.lean +++ b/Mathlib/AlgebraicGeometry/Spec.lean @@ -196,7 +196,7 @@ theorem Spec.basicOpen_hom_ext {X : RingedSpace.{u}} {R : CommRingCat.{u}} specialize h r simp only [sheafedSpaceObj_carrier, Functor.op_obj, unop_op, TopCat.Presheaf.pushforwardObj_obj, sheafedSpaceObj_presheaf, Category.assoc] at h - rw [NatTrans.comp_app, ←h] + rw [NatTrans.comp_app, ← h] congr simp set_option linter.uppercaseLean3 false in @@ -450,7 +450,7 @@ instance isLocalizedModule_toPushforwardStalkAlgHom : exact (IsLocalization.map_units ((structureSheaf R).presheaf.stalk p) ⟨x, hx⟩).map _ · apply isLocalizedModule_toPushforwardStalkAlgHom_aux · intro x hx - rw [toPushforwardStalkAlgHom_apply, ←(toPushforwardStalk (algebraMap R S) p).map_zero, + rw [toPushforwardStalkAlgHom_apply, ← (toPushforwardStalk (algebraMap R S) p).map_zero, toPushforwardStalk] at hx -- Porting note : this `change` is manually rewriting `comp_apply` change _ = (TopCat.Presheaf.germ (Spec.topMap (algebraMap ↑R ↑S) _* (structureSheaf ↑S).val) diff --git a/Mathlib/AlgebraicTopology/FundamentalGroupoid/Basic.lean b/Mathlib/AlgebraicTopology/FundamentalGroupoid/Basic.lean index fa23b34bd1634..14cacfa13b469 100644 --- a/Mathlib/AlgebraicTopology/FundamentalGroupoid/Basic.lean +++ b/Mathlib/AlgebraicTopology/FundamentalGroupoid/Basic.lean @@ -114,7 +114,7 @@ def reflTransSymm (p : Path x₀ x₁) : Homotopy (Path.refl x₀) (p.trans p.sy /-- For any path `p` from `x₀` to `x₁`, we have a homotopy from the constant path based at `x₁` to `p.symm.trans p`. -/ def reflSymmTrans (p : Path x₀ x₁) : Homotopy (Path.refl x₁) (p.symm.trans p) := - (reflTransSymm p.symm).cast rfl <| congr_arg _ Path.symm_symm + (reflTransSymm p.symm).cast rfl <| congr_arg _ (Path.symm_symm _) #align path.homotopy.refl_symm_trans Path.Homotopy.reflSymmTrans end diff --git a/Mathlib/Analysis/Analytic/Constructions.lean b/Mathlib/Analysis/Analytic/Constructions.lean index 2e32092155058..50f1e7121ec2e 100644 --- a/Mathlib/Analysis/Analytic/Constructions.lean +++ b/Mathlib/Analysis/Analytic/Constructions.lean @@ -60,7 +60,7 @@ lemma FormalMultilinearSeries.radius_prod_eq_min have := ((p.isLittleO_one_of_lt_radius hr.1).add (q.isLittleO_one_of_lt_radius hr.2)).isBigO refine (p.prod q).le_radius_of_isBigO ((isBigO_of_le _ λ n ↦ ?_).trans this) - rw [norm_mul, norm_norm, ←add_mul, norm_mul] + rw [norm_mul, norm_norm, ← add_mul, norm_mul] refine mul_le_mul_of_nonneg_right ?_ (norm_nonneg _) rw [FormalMultilinearSeries.prod, ContinuousMultilinearMap.op_norm_prod] refine (max_le_add_of_nonneg (norm_nonneg _) (norm_nonneg _)).trans ?_ @@ -149,7 +149,7 @@ lemma analyticAt_smul [NormedSpace 𝕝 E] [IsScalarTower 𝕜 𝕝 E] (z : 𝕝 AnalyticAt 𝕜 (fun x : 𝕝 × E ↦ x.1 • x.2) z := (ContinuousLinearMap.lsmul 𝕜 𝕝).analyticAt_bilinear z -/-- Multiplication in a normed algebra over `𝕜` is -/ +/-- Multiplication in a normed algebra over `𝕜` is analytic. -/ lemma analyticAt_mul (z : A × A) : AnalyticAt 𝕜 (fun x : A × A ↦ x.1 * x.2) z := (ContinuousLinearMap.mul 𝕜 A).analyticAt_bilinear z @@ -165,12 +165,12 @@ lemma AnalyticOn.smul [NormedSpace 𝕝 F] [IsScalarTower 𝕜 𝕝 F] {f : E AnalyticOn 𝕜 (fun x ↦ f x • g x) s := fun _ m ↦ (hf _ m).smul (hg _ m) -/-- Multiplication of analytic functions (valued in a normd `𝕜`-algebra) is analytic. -/ +/-- Multiplication of analytic functions (valued in a normed `𝕜`-algebra) is analytic. -/ lemma AnalyticAt.mul {f g : E → A} {z : E} (hf : AnalyticAt 𝕜 f z) (hg : AnalyticAt 𝕜 g z) : AnalyticAt 𝕜 (fun x ↦ f x * g x) z := (analyticAt_mul _).comp₂ hf hg -/-- Multiplication of analytic functions (valued in a normd `𝕜`-algebra) is analytic. -/ +/-- Multiplication of analytic functions (valued in a normed `𝕜`-algebra) is analytic. -/ lemma AnalyticOn.mul {f g : E → A} {s : Set E} (hf : AnalyticOn 𝕜 f s) (hg : AnalyticOn 𝕜 g s) : AnalyticOn 𝕜 (fun x ↦ f x * g x) s := fun _ m ↦ (hf _ m).mul (hg _ m) @@ -178,10 +178,12 @@ lemma AnalyticOn.mul {f g : E → A} {s : Set E} (hf : AnalyticOn 𝕜 f s) (hg /-- Powers of analytic functions (into a normed `𝕜`-algebra) are analytic. -/ lemma AnalyticAt.pow {f : E → A} {z : E} (hf : AnalyticAt 𝕜 f z) (n : ℕ) : AnalyticAt 𝕜 (fun x ↦ f x ^ n) z := by - induction' n with m hm - · simp only [Nat.zero_eq, pow_zero] + induction n with + | zero => + simp only [Nat.zero_eq, pow_zero] apply analyticAt_const - · simp only [pow_succ] + | succ m hm => + simp only [pow_succ] exact hf.mul hm /-- Powers of analytic functions (into a normed `𝕜`-algebra) are analytic. -/ @@ -209,7 +211,7 @@ lemma formalMultilinearSeries_geometric_radius (𝕜) [NontriviallyNormedField (formalMultilinearSeries_geometric 𝕜 A).radius = 1 := by apply le_antisymm · refine le_of_forall_nnreal_lt (fun r hr ↦ ?_) - rw [←coe_one, ENNReal.coe_le_coe] + rw [← coe_one, ENNReal.coe_le_coe] have := FormalMultilinearSeries.isLittleO_one_of_lt_radius _ hr simp_rw [formalMultilinearSeries_geometric_apply_norm, one_mul] at this contrapose! this @@ -220,10 +222,10 @@ lemma formalMultilinearSeries_geometric_radius (𝕜) [NontriviallyNormedField intro n hn push_neg rwa [norm_pow, one_lt_pow_iff_of_nonneg (norm_nonneg _) hn, - Real.norm_of_nonneg (NNReal.coe_nonneg _), ←NNReal.coe_one, + Real.norm_of_nonneg (NNReal.coe_nonneg _), ← NNReal.coe_one, NNReal.coe_lt_coe] · refine le_of_forall_nnreal_lt (fun r hr ↦ ?_) - rw [←Nat.cast_one, ENNReal.coe_lt_coe_nat, Nat.cast_one] at hr + rw [← Nat.cast_one, ENNReal.coe_lt_coe_nat, Nat.cast_one] at hr apply FormalMultilinearSeries.le_radius_of_isBigO simp_rw [formalMultilinearSeries_geometric_apply_norm, one_mul] refine isBigO_of_le atTop (fun n ↦ ?_) @@ -242,7 +244,7 @@ lemma hasFPowerSeriesOnBall_inv_one_sub List.prod_ofFn, Finset.prod_const, Finset.card_univ, Fintype.card_fin] apply hasSum_geometric_of_norm_lt_1 - simpa only [←ofReal_one, Metric.emetric_ball, Metric.ball, + simpa only [← ofReal_one, Metric.emetric_ball, Metric.ball, dist_eq_norm, sub_zero] using hy lemma analyticAt_inv_one_sub (𝕝 : Type*) [NontriviallyNormedField 𝕝] [NormedAlgebra 𝕜 𝕝] : diff --git a/Mathlib/Analysis/Analytic/RadiusLiminf.lean b/Mathlib/Analysis/Analytic/RadiusLiminf.lean index 97fe3033bab65..4c62a7311a5e4 100644 --- a/Mathlib/Analysis/Analytic/RadiusLiminf.lean +++ b/Mathlib/Analysis/Analytic/RadiusLiminf.lean @@ -51,7 +51,7 @@ theorem radius_eq_liminf : obtain ⟨a, ha, H⟩ := this apply le_liminf_of_le · infer_param - · rw [←eventually_map] + · rw [← eventually_map] refine' H.mp ((eventually_gt_atTop 0).mono fun n hn₀ hn => (this _ hn₀).2 (NNReal.coe_le_coe.1 _)) push_cast diff --git a/Mathlib/Analysis/Asymptotics/Asymptotics.lean b/Mathlib/Analysis/Asymptotics/Asymptotics.lean index ff996db177b2d..4a4494c843e03 100644 --- a/Mathlib/Analysis/Asymptotics/Asymptotics.lean +++ b/Mathlib/Analysis/Asymptotics/Asymptotics.lean @@ -150,7 +150,7 @@ theorem isBigO_iff'' {g : α → E'''} : obtain ⟨c, ⟨hc_pos, hc⟩⟩ := h refine ⟨c⁻¹, ⟨by positivity, ?_⟩⟩ filter_upwards [hc] with x hx - rwa [←inv_inv c, inv_mul_le_iff (by positivity)] at hx + rwa [← inv_inv c, inv_mul_le_iff (by positivity)] at hx theorem IsBigO.of_bound (c : ℝ) (h : ∀ᶠ x in l, ‖f x‖ ≤ c * ‖g x‖) : f =O[l] g := isBigO_iff.2 ⟨c, h⟩ @@ -1692,7 +1692,7 @@ theorem IsLittleO.pow {f : α → R} {g : α → 𝕜} (h : f =o[l] g) {n : ℕ} (fun x => f x ^ n) =o[l] fun x => g x ^ n := by obtain ⟨n, rfl⟩ := Nat.exists_eq_succ_of_ne_zero hn.ne'; clear hn induction' n with n ihn - · simpa only [Nat.zero_eq, ←Nat.one_eq_succ_zero, pow_one] + · simpa only [Nat.zero_eq, ← Nat.one_eq_succ_zero, pow_one] · convert h.mul ihn <;> simp [pow_succ] #align asymptotics.is_o.pow Asymptotics.IsLittleO.pow diff --git a/Mathlib/Analysis/BoxIntegral/Partition/Additive.lean b/Mathlib/Analysis/BoxIntegral/Partition/Additive.lean index 70895c4c645ce..9781e5d855bbe 100644 --- a/Mathlib/Analysis/BoxIntegral/Partition/Additive.lean +++ b/Mathlib/Analysis/BoxIntegral/Partition/Additive.lean @@ -145,7 +145,7 @@ map. -/ @[simps (config := .asFn)] def map (f : ι →ᵇᵃ[I₀] M) (g : M →+ N) : ι →ᵇᵃ[I₀] N where toFun := g ∘ f - sum_partition_boxes' I hI π hπ := by simp_rw [comp, ← g.map_sum, f.sum_partition_boxes hI hπ] + sum_partition_boxes' I hI π hπ := by simp_rw [comp, ← map_sum, f.sum_partition_boxes hI hπ] #align box_integral.box_additive_map.map BoxIntegral.BoxAdditiveMap.map /-- If `f` is a box additive function on subboxes of `I` and `π₁`, `π₂` are two prepartitions of diff --git a/Mathlib/Analysis/Calculus/ContDiff/Basic.lean b/Mathlib/Analysis/Calculus/ContDiff/Basic.lean index 747b4a7f80104..6fc432e534efd 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/Basic.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/Basic.lean @@ -1833,7 +1833,7 @@ theorem LocalHomeomorph.contDiffAt_symm [CompleteSpace E] (f : LocalHomeomorph E obtain ⟨t, htu, ht, htf⟩ := mem_nhds_iff.mp (Filter.inter_mem hu h_nhds) use f.target ∩ f.symm ⁻¹' t refine' ⟨IsOpen.mem_nhds _ _, _⟩ - · exact f.preimage_open_of_open_symm ht + · exact f.isOpen_inter_preimage_symm ht · exact mem_inter ha (mem_preimage.mpr htf) intro x hx obtain ⟨hxu, e, he⟩ := htu hx.2 @@ -1895,6 +1895,33 @@ theorem Homeomorph.contDiff_symm_deriv [CompleteSpace 𝕜] (f : 𝕜 ≃ₜ f.toLocalHomeomorph.contDiffAt_symm_deriv (h₀ _) (mem_univ x) (hf' _) hf.contDiffAt #align homeomorph.cont_diff_symm_deriv Homeomorph.contDiff_symm_deriv +namespace LocalHomeomorph + +variable (𝕜) + +/-- Restrict a local homeomorphism to the subsets of the source and target +that consist of points `x ∈ f.source`, `y = f x ∈ f.target` +such that `f` is `C^n` at `x` and `f.symm` is `C^n` at `y`. + +Note that `n` is a natural number, not `∞`, +because the set of points of `C^∞`-smoothness of `f` is not guaranteed to be open. -/ +@[simps! apply symm_apply source target] +def restrContDiff (f : LocalHomeomorph E F) (n : ℕ) : LocalHomeomorph E F := + haveI H : f.IsImage {x | ContDiffAt 𝕜 n f x ∧ ContDiffAt 𝕜 n f.symm (f x)} + {y | ContDiffAt 𝕜 n f.symm y ∧ ContDiffAt 𝕜 n f (f.symm y)} := fun x hx ↦ by + simp [hx, and_comm] + H.restr <| isOpen_iff_mem_nhds.2 <| fun x ⟨hxs, hxf, hxf'⟩ ↦ + inter_mem (f.open_source.mem_nhds hxs) <| hxf.eventually.and <| + f.continuousAt hxs hxf'.eventually + +lemma contDiffOn_restrContDiff_source (f : LocalHomeomorph E F) (n : ℕ) : + ContDiffOn 𝕜 n f (f.restrContDiff 𝕜 n).source := fun _x hx ↦ hx.2.1.contDiffWithinAt + +lemma contDiffOn_restrContDiff_target (f : LocalHomeomorph E F) (n : ℕ) : + ContDiffOn 𝕜 n f.symm (f.restrContDiff 𝕜 n).target := fun _x hx ↦ hx.2.1.contDiffWithinAt + +end LocalHomeomorph + end FunctionInverse section deriv @@ -1937,11 +1964,11 @@ theorem contDiffOn_succ_iff_derivWithin {n : ℕ} (hs : UniqueDiffOn 𝕜 s₂) /-- A function is `C^(n + 1)` on an open domain if and only if it is differentiable there, and its derivative (formulated with `deriv`) is `C^n`. -/ -theorem contDiffOn_succ_iff_deriv_of_open {n : ℕ} (hs : IsOpen s₂) : +theorem contDiffOn_succ_iff_deriv_of_isOpen {n : ℕ} (hs : IsOpen s₂) : ContDiffOn 𝕜 (n + 1 : ℕ) f₂ s₂ ↔ DifferentiableOn 𝕜 f₂ s₂ ∧ ContDiffOn 𝕜 n (deriv f₂) s₂ := by rw [contDiffOn_succ_iff_derivWithin hs.uniqueDiffOn] - exact Iff.rfl.and (contDiffOn_congr fun _ => derivWithin_of_open hs) -#align cont_diff_on_succ_iff_deriv_of_open contDiffOn_succ_iff_deriv_of_open + exact Iff.rfl.and (contDiffOn_congr fun _ => derivWithin_of_isOpen hs) +#align cont_diff_on_succ_iff_deriv_of_open contDiffOn_succ_iff_deriv_of_isOpen /-- A function is `C^∞` on a domain with unique derivatives if and only if it is differentiable there, and its derivative (formulated with `derivWithin`) is `C^∞`. -/ @@ -1961,11 +1988,11 @@ theorem contDiffOn_top_iff_derivWithin (hs : UniqueDiffOn 𝕜 s₂) : /-- A function is `C^∞` on an open domain if and only if it is differentiable there, and its derivative (formulated with `deriv`) is `C^∞`. -/ -theorem contDiffOn_top_iff_deriv_of_open (hs : IsOpen s₂) : +theorem contDiffOn_top_iff_deriv_of_isOpen (hs : IsOpen s₂) : ContDiffOn 𝕜 ∞ f₂ s₂ ↔ DifferentiableOn 𝕜 f₂ s₂ ∧ ContDiffOn 𝕜 ∞ (deriv f₂) s₂ := by rw [contDiffOn_top_iff_derivWithin hs.uniqueDiffOn] - exact Iff.rfl.and <| contDiffOn_congr fun _ => derivWithin_of_open hs -#align cont_diff_on_top_iff_deriv_of_open contDiffOn_top_iff_deriv_of_open + exact Iff.rfl.and <| contDiffOn_congr fun _ => derivWithin_of_isOpen hs +#align cont_diff_on_top_iff_deriv_of_open contDiffOn_top_iff_deriv_of_isOpen protected theorem ContDiffOn.derivWithin (hf : ContDiffOn 𝕜 n f₂ s₂) (hs : UniqueDiffOn 𝕜 s₂) (hmn : m + 1 ≤ n) : ContDiffOn 𝕜 m (derivWithin f₂ s₂) s₂ := by @@ -1978,26 +2005,26 @@ protected theorem ContDiffOn.derivWithin (hf : ContDiffOn 𝕜 n f₂ s₂) (hs exact ((contDiffOn_succ_iff_derivWithin hs).1 (hf.of_le hmn)).2 #align cont_diff_on.deriv_within ContDiffOn.derivWithin -theorem ContDiffOn.deriv_of_open (hf : ContDiffOn 𝕜 n f₂ s₂) (hs : IsOpen s₂) (hmn : m + 1 ≤ n) : +theorem ContDiffOn.deriv_of_isOpen (hf : ContDiffOn 𝕜 n f₂ s₂) (hs : IsOpen s₂) (hmn : m + 1 ≤ n) : ContDiffOn 𝕜 m (deriv f₂) s₂ := - (hf.derivWithin hs.uniqueDiffOn hmn).congr fun _ hx => (derivWithin_of_open hs hx).symm -#align cont_diff_on.deriv_of_open ContDiffOn.deriv_of_open + (hf.derivWithin hs.uniqueDiffOn hmn).congr fun _ hx => (derivWithin_of_isOpen hs hx).symm +#align cont_diff_on.deriv_of_open ContDiffOn.deriv_of_isOpen theorem ContDiffOn.continuousOn_derivWithin (h : ContDiffOn 𝕜 n f₂ s₂) (hs : UniqueDiffOn 𝕜 s₂) (hn : 1 ≤ n) : ContinuousOn (derivWithin f₂ s₂) s₂ := ((contDiffOn_succ_iff_derivWithin hs).1 (h.of_le hn)).2.continuousOn #align cont_diff_on.continuous_on_deriv_within ContDiffOn.continuousOn_derivWithin -theorem ContDiffOn.continuousOn_deriv_of_open (h : ContDiffOn 𝕜 n f₂ s₂) (hs : IsOpen s₂) +theorem ContDiffOn.continuousOn_deriv_of_isOpen (h : ContDiffOn 𝕜 n f₂ s₂) (hs : IsOpen s₂) (hn : 1 ≤ n) : ContinuousOn (deriv f₂) s₂ := - ((contDiffOn_succ_iff_deriv_of_open hs).1 (h.of_le hn)).2.continuousOn -#align cont_diff_on.continuous_on_deriv_of_open ContDiffOn.continuousOn_deriv_of_open + ((contDiffOn_succ_iff_deriv_of_isOpen hs).1 (h.of_le hn)).2.continuousOn +#align cont_diff_on.continuous_on_deriv_of_open ContDiffOn.continuousOn_deriv_of_isOpen /-- A function is `C^(n + 1)` if and only if it is differentiable, and its derivative (formulated in terms of `deriv`) is `C^n`. -/ theorem contDiff_succ_iff_deriv {n : ℕ} : ContDiff 𝕜 (n + 1 : ℕ) f₂ ↔ Differentiable 𝕜 f₂ ∧ ContDiff 𝕜 n (deriv f₂) := by - simp only [← contDiffOn_univ, contDiffOn_succ_iff_deriv_of_open, isOpen_univ, + simp only [← contDiffOn_univ, contDiffOn_succ_iff_deriv_of_isOpen, isOpen_univ, differentiableOn_univ] #align cont_diff_succ_iff_deriv contDiff_succ_iff_deriv diff --git a/Mathlib/Analysis/Calculus/ContDiff/Defs.lean b/Mathlib/Analysis/Calculus/ContDiff/Defs.lean index 9ff8127679f5d..837554e04ecad 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/Defs.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/Defs.lean @@ -1156,12 +1156,12 @@ theorem contDiffOn_succ_iff_has_fderiv_within {n : ℕ} (hs : UniqueDiffOn 𝕜 /-- A function is `C^(n + 1)` on an open domain if and only if it is differentiable there, and its derivative (expressed with `fderiv`) is `C^n`. -/ -theorem contDiffOn_succ_iff_fderiv_of_open {n : ℕ} (hs : IsOpen s) : +theorem contDiffOn_succ_iff_fderiv_of_isOpen {n : ℕ} (hs : IsOpen s) : ContDiffOn 𝕜 (n + 1 : ℕ) f s ↔ DifferentiableOn 𝕜 f s ∧ ContDiffOn 𝕜 n (fun y => fderiv 𝕜 f y) s := by rw [contDiffOn_succ_iff_fderivWithin hs.uniqueDiffOn] exact Iff.rfl.and (contDiffOn_congr fun x hx ↦ fderivWithin_of_isOpen hs hx) -#align cont_diff_on_succ_iff_fderiv_of_open contDiffOn_succ_iff_fderiv_of_open +#align cont_diff_on_succ_iff_fderiv_of_open contDiffOn_succ_iff_fderiv_of_isOpen /-- A function is `C^∞` on a domain with unique derivatives if and only if it is differentiable there, and its derivative (expressed with `fderivWithin`) is `C^∞`. -/ @@ -1182,11 +1182,11 @@ theorem contDiffOn_top_iff_fderivWithin (hs : UniqueDiffOn 𝕜 s) : /-- A function is `C^∞` on an open domain if and only if it is differentiable there, and its derivative (expressed with `fderiv`) is `C^∞`. -/ -theorem contDiffOn_top_iff_fderiv_of_open (hs : IsOpen s) : +theorem contDiffOn_top_iff_fderiv_of_isOpen (hs : IsOpen s) : ContDiffOn 𝕜 ∞ f s ↔ DifferentiableOn 𝕜 f s ∧ ContDiffOn 𝕜 ∞ (fun y => fderiv 𝕜 f y) s := by rw [contDiffOn_top_iff_fderivWithin hs.uniqueDiffOn] exact Iff.rfl.and <| contDiffOn_congr fun x hx ↦ fderivWithin_of_isOpen hs hx -#align cont_diff_on_top_iff_fderiv_of_open contDiffOn_top_iff_fderiv_of_open +#align cont_diff_on_top_iff_fderiv_of_open contDiffOn_top_iff_fderiv_of_isOpen protected theorem ContDiffOn.fderivWithin (hf : ContDiffOn 𝕜 n f s) (hs : UniqueDiffOn 𝕜 s) (hmn : m + 1 ≤ n) : ContDiffOn 𝕜 m (fun y => fderivWithin 𝕜 f s y) s := by @@ -1199,20 +1199,20 @@ protected theorem ContDiffOn.fderivWithin (hf : ContDiffOn 𝕜 n f s) (hs : Uni exact ((contDiffOn_succ_iff_fderivWithin hs).1 (hf.of_le hmn)).2 #align cont_diff_on.fderiv_within ContDiffOn.fderivWithin -theorem ContDiffOn.fderiv_of_open (hf : ContDiffOn 𝕜 n f s) (hs : IsOpen s) (hmn : m + 1 ≤ n) : +theorem ContDiffOn.fderiv_of_isOpen (hf : ContDiffOn 𝕜 n f s) (hs : IsOpen s) (hmn : m + 1 ≤ n) : ContDiffOn 𝕜 m (fun y => fderiv 𝕜 f y) s := (hf.fderivWithin hs.uniqueDiffOn hmn).congr fun _ hx => (fderivWithin_of_isOpen hs hx).symm -#align cont_diff_on.fderiv_of_open ContDiffOn.fderiv_of_open +#align cont_diff_on.fderiv_of_open ContDiffOn.fderiv_of_isOpen theorem ContDiffOn.continuousOn_fderivWithin (h : ContDiffOn 𝕜 n f s) (hs : UniqueDiffOn 𝕜 s) (hn : 1 ≤ n) : ContinuousOn (fun x => fderivWithin 𝕜 f s x) s := ((contDiffOn_succ_iff_fderivWithin hs).1 (h.of_le hn)).2.continuousOn #align cont_diff_on.continuous_on_fderiv_within ContDiffOn.continuousOn_fderivWithin -theorem ContDiffOn.continuousOn_fderiv_of_open (h : ContDiffOn 𝕜 n f s) (hs : IsOpen s) +theorem ContDiffOn.continuousOn_fderiv_of_isOpen (h : ContDiffOn 𝕜 n f s) (hs : IsOpen s) (hn : 1 ≤ n) : ContinuousOn (fun x => fderiv 𝕜 f x) s := - ((contDiffOn_succ_iff_fderiv_of_open hs).1 (h.of_le hn)).2.continuousOn -#align cont_diff_on.continuous_on_fderiv_of_open ContDiffOn.continuousOn_fderiv_of_open + ((contDiffOn_succ_iff_fderiv_of_isOpen hs).1 (h.of_le hn)).2.continuousOn +#align cont_diff_on.continuous_on_fderiv_of_open ContDiffOn.continuousOn_fderiv_of_isOpen /-! ### Functions with a Taylor series on the whole space -/ diff --git a/Mathlib/Analysis/Calculus/Deriv/Basic.lean b/Mathlib/Analysis/Calculus/Deriv/Basic.lean index 9f571c7f8536e..3b25240ada6f8 100644 --- a/Mathlib/Analysis/Calculus/Deriv/Basic.lean +++ b/Mathlib/Analysis/Calculus/Deriv/Basic.lean @@ -531,9 +531,9 @@ theorem derivWithin_inter (ht : t ∈ 𝓝 x) : derivWithin f (s ∩ t) x = deri theorem derivWithin_of_mem_nhds (h : s ∈ 𝓝 x) : derivWithin f s x = deriv f x := by simp only [derivWithin, deriv, fderivWithin_of_mem_nhds h] -theorem derivWithin_of_open (hs : IsOpen s) (hx : x ∈ s) : derivWithin f s x = deriv f x := +theorem derivWithin_of_isOpen (hs : IsOpen s) (hx : x ∈ s) : derivWithin f s x = deriv f x := derivWithin_of_mem_nhds (hs.mem_nhds hx) -#align deriv_within_of_open derivWithin_of_open +#align deriv_within_of_open derivWithin_of_isOpen theorem deriv_mem_iff {f : 𝕜 → F} {s : Set F} {x : 𝕜} : deriv f x ∈ s ↔ diff --git a/Mathlib/Analysis/Calculus/FormalMultilinearSeries.lean b/Mathlib/Analysis/Calculus/FormalMultilinearSeries.lean index bd47ced3665f5..92956f56c4484 100644 --- a/Mathlib/Analysis/Calculus/FormalMultilinearSeries.lean +++ b/Mathlib/Analysis/Calculus/FormalMultilinearSeries.lean @@ -3,7 +3,7 @@ Copyright (c) 2019 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ -import Mathlib.Analysis.NormedSpace.MultilinearCurrying +import Mathlib.Analysis.NormedSpace.Multilinear.Curry #align_import analysis.calculus.formal_multilinear_series from "leanprover-community/mathlib"@"f2ce6086713c78a7f880485f7917ea547a215982" diff --git a/Mathlib/Analysis/Calculus/LineDeriv/Basic.lean b/Mathlib/Analysis/Calculus/LineDeriv/Basic.lean index c552f2941eb74..57ca9ef4e651b 100644 --- a/Mathlib/Analysis/Calculus/LineDeriv/Basic.lean +++ b/Mathlib/Analysis/Calculus/LineDeriv/Basic.lean @@ -278,7 +278,7 @@ theorem lineDerivWithin_of_mem_nhds (h : s ∈ 𝓝 x) : apply (Continuous.continuousAt _).preimage_mem_nhds (by simpa using h) continuity -theorem lineDerivWithin_of_open (hs : IsOpen s) (hx : x ∈ s) : +theorem lineDerivWithin_of_isOpen (hs : IsOpen s) (hx : x ∈ s) : lineDerivWithin 𝕜 f s x v = lineDeriv 𝕜 f x v := lineDerivWithin_of_mem_nhds (hs.mem_nhds hx) diff --git a/Mathlib/Analysis/Calculus/Rademacher.lean b/Mathlib/Analysis/Calculus/Rademacher.lean index c545b07de01ba..211f345769d31 100644 --- a/Mathlib/Analysis/Calculus/Rademacher.lean +++ b/Mathlib/Analysis/Calculus/Rademacher.lean @@ -220,7 +220,7 @@ theorem ae_lineDeriv_sum_eq ContDiff.lipschitzWith_of_hasCompactSupport g_comp g_smooth le_top simp_rw [integral_lineDeriv_mul_eq hf g_lip g_comp] simp_rw [(g_smooth.differentiable le_top).differentiableAt.lineDeriv_eq_fderiv] - simp only [map_neg, map_sum, SMulHomClass.map_smul, smul_eq_mul, neg_mul] + simp only [map_neg, _root_.map_sum, SMulHomClass.map_smul, smul_eq_mul, neg_mul] simp only [integral_neg, mul_neg, Finset.sum_neg_distrib, neg_inj] exact S2 suffices B : ∀ i ∈ s, Integrable (fun x ↦ a i * (fderiv ℝ g x (v i) * f x)) μ by diff --git a/Mathlib/Analysis/Complex/AbsMax.lean b/Mathlib/Analysis/Complex/AbsMax.lean index 94d6aa004782f..9d6569f40c745 100644 --- a/Mathlib/Analysis/Complex/AbsMax.lean +++ b/Mathlib/Analysis/Complex/AbsMax.lean @@ -240,7 +240,7 @@ theorem norm_eqOn_of_isPreconnected_of_isMaxOn {f : E → F} {U : Set E} {c : E} using isOpen_setOf_mem_nhds_and_isMaxOn_norm hd have hVne : (U ∩ V).Nonempty := ⟨c, hcU, hcU, hm⟩ set W := U ∩ {z | ‖f z‖ ≠ ‖f c‖} - have hWo : IsOpen W := hd.continuousOn.norm.preimage_open_of_open ho isOpen_ne + have hWo : IsOpen W := hd.continuousOn.norm.isOpen_inter_preimage ho isOpen_ne have hdVW : Disjoint V W := disjoint_left.mpr fun x hxV hxW => hxW.2 (hV x hxV) have hUVW : U ⊆ V ∪ W := fun x hx => (eq_or_ne ‖f x‖ ‖f c‖).imp (fun h => ⟨hx, fun y hy => (hm hy).out.trans_eq h.symm⟩) diff --git a/Mathlib/Analysis/Complex/Polynomial.lean b/Mathlib/Analysis/Complex/Polynomial.lean index 1652ce99a70f5..637bf597f822f 100644 --- a/Mathlib/Analysis/Complex/Polynomial.lean +++ b/Mathlib/Analysis/Complex/Polynomial.lean @@ -27,7 +27,7 @@ namespace Complex /-- **Fundamental theorem of algebra**: every non constant complex polynomial has a root -/ theorem exists_root {f : ℂ[X]} (hf : 0 < degree f) : ∃ z : ℂ, IsRoot f z := by - by_contra' hf' + by_contra! hf' /- Since `f` has no roots, `f⁻¹` is differentiable. And since `f` is a polynomial, it tends to infinity at infinity, thus `f⁻¹` tends to zero at infinity. By Liouville's theorem, `f⁻¹ = 0`. -/ have (z : ℂ) : (f.eval z)⁻¹ = 0 := diff --git a/Mathlib/Analysis/ConstantSpeed.lean b/Mathlib/Analysis/ConstantSpeed.lean index 46d6c05471fc5..9732aca5e18ba 100644 --- a/Mathlib/Analysis/ConstantSpeed.lean +++ b/Mathlib/Analysis/ConstantSpeed.lean @@ -159,7 +159,7 @@ theorem hasConstantSpeedOnWith_zero_iff : dsimp [HasConstantSpeedOnWith] simp only [zero_mul, ENNReal.ofReal_zero, ← eVariationOn.eq_zero_iff] constructor - · by_contra' + · by_contra! obtain ⟨h, hfs⟩ := this simp_rw [ne_eq, eVariationOn.eq_zero_iff] at hfs h push_neg at hfs diff --git a/Mathlib/Analysis/Convex/Combination.lean b/Mathlib/Analysis/Convex/Combination.lean index f8671900bf50b..fdbee74f4f2bb 100644 --- a/Mathlib/Analysis/Convex/Combination.lean +++ b/Mathlib/Analysis/Convex/Combination.lean @@ -76,7 +76,7 @@ theorem Finset.centerMass_singleton (hw : w i ≠ 0) : ({i} : Finset ι).centerM lemma Finset.centerMass_smul_left {c : R'} [Module R' R] [Module R' E] [SMulCommClass R' R R] [IsScalarTower R' R R] [SMulCommClass R R' E] [IsScalarTower R' R E] (hc : c ≠ 0) : t.centerMass (c • w) z = t.centerMass w z := by - simp [centerMass, -smul_assoc, smul_assoc c, ←smul_sum, smul_inv₀, smul_smul_smul_comm, hc] + simp [centerMass, -smul_assoc, smul_assoc c, ← smul_sum, smul_inv₀, smul_smul_smul_comm, hc] theorem Finset.centerMass_eq_of_sum_1 (hw : ∑ i in t, w i = 1) : t.centerMass w z = ∑ i in t, w i • z i := by @@ -160,7 +160,7 @@ variable {z} lemma Finset.centerMass_of_sum_add_sum_eq_zero {s t : Finset ι} (hw : ∑ i in s, w i + ∑ i in t, w i = 0) (hz : ∑ i in s, w i • z i + ∑ i in t, w i • z i = 0) : s.centerMass w z = t.centerMass w z := by - simp [centerMass, eq_neg_of_add_eq_zero_right hw, eq_neg_of_add_eq_zero_left hz, ←neg_inv] + simp [centerMass, eq_neg_of_add_eq_zero_right hw, eq_neg_of_add_eq_zero_left hz, ← neg_inv] /-- The center of mass of a finite subset of a convex set belongs to the set provided that all weights are non-negative, and the total weight is positive. -/ @@ -240,7 +240,7 @@ theorem Finset.centerMass_mem_convexHull (t : Finset ι) {w : ι → R} (hw₀ : /-- A version of `Finset.centerMass_mem_convexHull` for when the weights are nonpositive. -/ lemma Finset.centerMass_mem_convexHull_of_nonpos (t : Finset ι) (hw₀ : ∀ i ∈ t, w i ≤ 0) (hws : ∑ i in t, w i < 0) (hz : ∀ i ∈ t, z i ∈ s) : t.centerMass w z ∈ convexHull R s := by - rw [←centerMass_neg_left] + rw [← centerMass_neg_left] exact Finset.centerMass_mem_convexHull _ (λ _i hi ↦ neg_nonneg.2 $ hw₀ _ hi) (by simpa) hz /-- A refinement of `Finset.centerMass_mem_convexHull` when the indexed family is a `Finset` of diff --git a/Mathlib/Analysis/Convex/Extreme.lean b/Mathlib/Analysis/Convex/Extreme.lean index ecb3290468a2b..66b12335c033e 100644 --- a/Mathlib/Analysis/Convex/Extreme.lean +++ b/Mathlib/Analysis/Convex/Extreme.lean @@ -164,7 +164,7 @@ theorem inter_extremePoints_subset_extremePoints_of_subset (hBA : B ⊆ A) : theorem IsExtreme.extremePoints_subset_extremePoints (hAB : IsExtreme 𝕜 A B) : B.extremePoints 𝕜 ⊆ A.extremePoints 𝕜 := - fun _ ↦ by simpa only [←isExtreme_singleton] using hAB.trans + fun _ ↦ by simpa only [← isExtreme_singleton] using hAB.trans #align is_extreme.extreme_points_subset_extreme_points IsExtreme.extremePoints_subset_extremePoints theorem IsExtreme.extremePoints_eq (hAB : IsExtreme 𝕜 A B) : @@ -273,7 +273,7 @@ theorem Convex.mem_extremePoints_iff_convex_diff (hA : Convex 𝕜 A) : rintro ⟨hxA, hAx⟩ refine' mem_extremePoints_iff_forall_segment.2 ⟨hxA, fun x₁ hx₁ x₂ hx₂ hx ↦ _⟩ rw [convex_iff_segment_subset] at hAx - by_contra' h + by_contra! h exact (hAx ⟨hx₁, fun hx₁ ↦ h.1 (mem_singleton_iff.2 hx₁)⟩ ⟨hx₂, fun hx₂ ↦ h.2 (mem_singleton_iff.2 hx₂)⟩ hx).2 rfl #align convex.mem_extreme_points_iff_convex_diff Convex.mem_extremePoints_iff_convex_diff diff --git a/Mathlib/Analysis/Convex/Function.lean b/Mathlib/Analysis/Convex/Function.lean index fdbb6368e40a7..a643e602b75f5 100644 --- a/Mathlib/Analysis/Convex/Function.lean +++ b/Mathlib/Analysis/Convex/Function.lean @@ -1093,9 +1093,9 @@ theorem OrderIso.strictConvexOn_symm (f : α ≃o β) (hf : StrictConcaveOn 𝕜 refine ⟨convex_univ, fun x _ y _ hxy a b ha hb hab => ?_⟩ obtain ⟨x', hx''⟩ := f.surjective.exists.mp ⟨x, rfl⟩ obtain ⟨y', hy''⟩ := f.surjective.exists.mp ⟨y, rfl⟩ - have hxy' : x' ≠ y' := by rw [←f.injective.ne_iff, ←hx'', ←hy'']; exact hxy + have hxy' : x' ≠ y' := by rw [← f.injective.ne_iff, ← hx'', ← hy'']; exact hxy simp only [hx'', hy'', OrderIso.symm_apply_apply, gt_iff_lt] - rw [←f.lt_iff_lt, OrderIso.apply_symm_apply] + rw [← f.lt_iff_lt, OrderIso.apply_symm_apply] exact hf.2 (by simp : x' ∈ univ) (by simp : y' ∈ univ) hxy' ha hb hab theorem OrderIso.convexOn_symm (f : α ≃o β) (hf : ConcaveOn 𝕜 univ f) : @@ -1104,7 +1104,7 @@ theorem OrderIso.convexOn_symm (f : α ≃o β) (hf : ConcaveOn 𝕜 univ f) : obtain ⟨x', hx''⟩ := f.surjective.exists.mp ⟨x, rfl⟩ obtain ⟨y', hy''⟩ := f.surjective.exists.mp ⟨y, rfl⟩ simp only [hx'', hy'', OrderIso.symm_apply_apply, gt_iff_lt] - rw [←f.le_iff_le, OrderIso.apply_symm_apply] + rw [← f.le_iff_le, OrderIso.apply_symm_apply] exact hf.2 (by simp : x' ∈ univ) (by simp : y' ∈ univ) ha hb hab theorem OrderIso.strictConcaveOn_symm (f : α ≃o β) (hf : StrictConvexOn 𝕜 univ f) : @@ -1112,9 +1112,9 @@ theorem OrderIso.strictConcaveOn_symm (f : α ≃o β) (hf : StrictConvexOn 𝕜 refine ⟨convex_univ, fun x _ y _ hxy a b ha hb hab => ?_⟩ obtain ⟨x', hx''⟩ := f.surjective.exists.mp ⟨x, rfl⟩ obtain ⟨y', hy''⟩ := f.surjective.exists.mp ⟨y, rfl⟩ - have hxy' : x' ≠ y' := by rw [←f.injective.ne_iff, ←hx'', ←hy'']; exact hxy + have hxy' : x' ≠ y' := by rw [← f.injective.ne_iff, ← hx'', ← hy'']; exact hxy simp only [hx'', hy'', OrderIso.symm_apply_apply, gt_iff_lt] - rw [←f.lt_iff_lt, OrderIso.apply_symm_apply] + rw [← f.lt_iff_lt, OrderIso.apply_symm_apply] exact hf.2 (by simp : x' ∈ univ) (by simp : y' ∈ univ) hxy' ha hb hab theorem OrderIso.concaveOn_symm (f : α ≃o β) (hf : ConvexOn 𝕜 univ f) : @@ -1123,7 +1123,7 @@ theorem OrderIso.concaveOn_symm (f : α ≃o β) (hf : ConvexOn 𝕜 univ f) : obtain ⟨x', hx''⟩ := f.surjective.exists.mp ⟨x, rfl⟩ obtain ⟨y', hy''⟩ := f.surjective.exists.mp ⟨y, rfl⟩ simp only [hx'', hy'', OrderIso.symm_apply_apply, gt_iff_lt] - rw [←f.le_iff_le, OrderIso.apply_symm_apply] + rw [← f.le_iff_le, OrderIso.apply_symm_apply] exact hf.2 (by simp : x' ∈ univ) (by simp : y' ∈ univ) ha hb hab end OrderIso @@ -1146,7 +1146,7 @@ lemma StrictConvexOn.eq_of_isMinOn (hf : StrictConvexOn 𝕜 s f) (hfx : IsMinOn calc f z < _ := hf.2 hx hy hxy (by norm_num) (by norm_num) $ by norm_num _ ≤ (2 : 𝕜)⁻¹ • f z + (2 : 𝕜)⁻¹ • f z := by gcongr; exacts [hfx hz, hfy hz] - _ = f z := by rw [←_root_.add_smul]; norm_num + _ = f z := by rw [← _root_.add_smul]; norm_num /-- A strictly concave function admits at most one global maximum. -/ lemma StrictConcaveOn.eq_of_isMaxOn (hf : StrictConcaveOn 𝕜 s f) (hfx : IsMaxOn f s x) diff --git a/Mathlib/Analysis/Convex/Gauge.lean b/Mathlib/Analysis/Convex/Gauge.lean index 5fc9e7888a611..93cdb3761c0cf 100644 --- a/Mathlib/Analysis/Convex/Gauge.lean +++ b/Mathlib/Analysis/Convex/Gauge.lean @@ -359,24 +359,24 @@ theorem interior_subset_gauge_lt_one (s : Set E) : interior s ⊆ { x | gauge s exact (gauge_le_of_mem hr₀.le hxr).trans_lt hr₁ #align interior_subset_gauge_lt_one interior_subset_gauge_lt_one -theorem gauge_lt_one_eq_self_of_open (hs₁ : Convex ℝ s) (hs₀ : (0 : E) ∈ s) (hs₂ : IsOpen s) : +theorem gauge_lt_one_eq_self_of_isOpen (hs₁ : Convex ℝ s) (hs₀ : (0 : E) ∈ s) (hs₂ : IsOpen s) : { x | gauge s x < 1 } = s := by refine' (gauge_lt_one_subset_self hs₁ ‹_› <| absorbent_nhds_zero <| hs₂.mem_nhds hs₀).antisymm _ convert interior_subset_gauge_lt_one s exact hs₂.interior_eq.symm -#align gauge_lt_one_eq_self_of_open gauge_lt_one_eq_self_of_open +#align gauge_lt_one_eq_self_of_open gauge_lt_one_eq_self_of_isOpen -- porting note: droped unneeded assumptions -theorem gauge_lt_one_of_mem_of_open (hs₂ : IsOpen s) {x : E} (hx : x ∈ s) : +theorem gauge_lt_one_of_mem_of_isOpen (hs₂ : IsOpen s) {x : E} (hx : x ∈ s) : gauge s x < 1 := interior_subset_gauge_lt_one s <| by rwa [hs₂.interior_eq] -#align gauge_lt_one_of_mem_of_open gauge_lt_one_of_mem_of_openₓ +#align gauge_lt_one_of_mem_of_open gauge_lt_one_of_mem_of_isOpenₓ -- porting note: droped unneeded assumptions theorem gauge_lt_of_mem_smul (x : E) (ε : ℝ) (hε : 0 < ε) (hs₂ : IsOpen s) (hx : x ∈ ε • s) : gauge s x < ε := by have : ε⁻¹ • x ∈ s := by rwa [← mem_smul_set_iff_inv_smul_mem₀ hε.ne'] - have h_gauge_lt := gauge_lt_one_of_mem_of_open hs₂ this + have h_gauge_lt := gauge_lt_one_of_mem_of_isOpen hs₂ this rwa [gauge_smul_of_nonneg (inv_nonneg.2 hε.le), smul_eq_mul, inv_mul_lt_iff hε, mul_one] at h_gauge_lt #align gauge_lt_of_mem_smul gauge_lt_of_mem_smulₓ @@ -480,14 +480,14 @@ def gaugeSeminorm (hs₀ : Balanced 𝕜 s) (hs₁ : Convex ℝ s) (hs₂ : Abso variable {hs₀ : Balanced 𝕜 s} {hs₁ : Convex ℝ s} {hs₂ : Absorbent ℝ s} [TopologicalSpace E] [ContinuousSMul ℝ E] -theorem gaugeSeminorm_lt_one_of_open (hs : IsOpen s) {x : E} (hx : x ∈ s) : +theorem gaugeSeminorm_lt_one_of_isOpen (hs : IsOpen s) {x : E} (hx : x ∈ s) : gaugeSeminorm hs₀ hs₁ hs₂ x < 1 := - gauge_lt_one_of_mem_of_open hs hx -#align gauge_seminorm_lt_one_of_open gaugeSeminorm_lt_one_of_open + gauge_lt_one_of_mem_of_isOpen hs hx +#align gauge_seminorm_lt_one_of_open gaugeSeminorm_lt_one_of_isOpen theorem gaugeSeminorm_ball_one (hs : IsOpen s) : (gaugeSeminorm hs₀ hs₁ hs₂).ball 0 1 = s := by rw [Seminorm.ball_zero_eq] - exact gauge_lt_one_eq_self_of_open hs₁ hs₂.zero_mem hs + exact gauge_lt_one_eq_self_of_isOpen hs₁ hs₂.zero_mem hs #align gauge_seminorm_ball_one gaugeSeminorm_ball_one end IsROrC diff --git a/Mathlib/Analysis/Convex/KreinMilman.lean b/Mathlib/Analysis/Convex/KreinMilman.lean index 1c880aa13fbfb..11da633722c85 100644 --- a/Mathlib/Analysis/Convex/KreinMilman.lean +++ b/Mathlib/Analysis/Convex/KreinMilman.lean @@ -120,5 +120,5 @@ lemma surjOn_extremePoints_image (f : E →A[ℝ] F) (hs : IsCompact s) : refine mem_image_of_mem _ ⟨hx, fun y hy z hz hxyz ↦ ?_⟩ have := by simpa using image_openSegment _ f.toAffineMap y z have := hw.2 (mem_image_of_mem _ hy) (mem_image_of_mem _ hz) $ by - rw [←this]; exact mem_image_of_mem _ hxyz + rw [← this]; exact mem_image_of_mem _ hxyz exact hyt ⟨hy, this.1⟩ ⟨hz, this.2⟩ hxyz diff --git a/Mathlib/Analysis/Convex/Mul.lean b/Mathlib/Analysis/Convex/Mul.lean index 10ef1adf91e27..da2720a806aa8 100644 --- a/Mathlib/Analysis/Convex/Mul.lean +++ b/Mathlib/Analysis/Convex/Mul.lean @@ -41,8 +41,8 @@ lemma ConvexOn.smul' (hf : ConvexOn 𝕜 s f) (hg : ConvexOn 𝕜 s g) (hf₀ : simp only [mul_add, add_smul, smul_add, mul_comm _ a]; abel _ = _ := by simp_rw [hab, mul_one] simp only [mul_add, add_smul, smul_add] - rw [←smul_smul_smul_comm a, ←smul_smul_smul_comm b, ←smul_smul_smul_comm a b, - ←smul_smul_smul_comm b b, smul_eq_mul, smul_eq_mul, smul_eq_mul, smul_eq_mul, mul_comm b, + rw [← smul_smul_smul_comm a, ← smul_smul_smul_comm b, ← smul_smul_smul_comm a b, + ← smul_smul_smul_comm b b, smul_eq_mul, smul_eq_mul, smul_eq_mul, smul_eq_mul, mul_comm b, add_comm _ ((b * b) • f y • g y), add_add_add_comm, add_comm ((a * b) • f y • g x)] lemma ConcaveOn.smul' (hf : ConcaveOn 𝕜 s f) (hg : ConcaveOn 𝕜 s g) (hf₀ : ∀ ⦃x⦄, x ∈ s → 0 ≤ f x) @@ -59,44 +59,44 @@ lemma ConcaveOn.smul' (hf : ConcaveOn 𝕜 s f) (hg : ConcaveOn 𝕜 s g) (hf₀ gcongr _ + (a * b) • ?_; exact hfg.smul_add_smul_le_smul_add_smul hx hy _ = _ := ?_ simp only [mul_add, add_smul, smul_add] - rw [←smul_smul_smul_comm a, ←smul_smul_smul_comm b, ←smul_smul_smul_comm a b, - ←smul_smul_smul_comm b b, smul_eq_mul, smul_eq_mul, smul_eq_mul, smul_eq_mul, mul_comm b a, + rw [← smul_smul_smul_comm a, ← smul_smul_smul_comm b, ← smul_smul_smul_comm a b, + ← smul_smul_smul_comm b b, smul_eq_mul, smul_eq_mul, smul_eq_mul, smul_eq_mul, mul_comm b a, add_comm ((a * b) • f x • g y), add_comm ((a * b) • f x • g y), add_add_add_comm] lemma ConvexOn.smul'' (hf : ConvexOn 𝕜 s f) (hg : ConvexOn 𝕜 s g) (hf₀ : ∀ ⦃x⦄, x ∈ s → f x ≤ 0) (hg₀ : ∀ ⦃x⦄, x ∈ s → g x ≤ 0) (hfg : AntivaryOn f g s) : ConcaveOn 𝕜 s (f • g) := by - rw [←neg_smul_neg] + rw [← neg_smul_neg] exact hf.neg.smul' hg.neg (fun x hx ↦ neg_nonneg.2 <| hf₀ hx) (fun x hx ↦ neg_nonneg.2 <| hg₀ hx) hfg.neg lemma ConcaveOn.smul'' (hf : ConcaveOn 𝕜 s f) (hg : ConcaveOn 𝕜 s g) (hf₀ : ∀ ⦃x⦄, x ∈ s → f x ≤ 0) (hg₀ : ∀ ⦃x⦄, x ∈ s → g x ≤ 0) (hfg : MonovaryOn f g s) : ConvexOn 𝕜 s (f • g) := by - rw [←neg_smul_neg] + rw [← neg_smul_neg] exact hf.neg.smul' hg.neg (fun x hx ↦ neg_nonneg.2 <| hf₀ hx) (fun x hx ↦ neg_nonneg.2 <| hg₀ hx) hfg.neg lemma ConvexOn.smul_concaveOn (hf : ConvexOn 𝕜 s f) (hg : ConcaveOn 𝕜 s g) (hf₀ : ∀ ⦃x⦄, x ∈ s → 0 ≤ f x) (hg₀ : ∀ ⦃x⦄, x ∈ s → g x ≤ 0) (hfg : AntivaryOn f g s) : ConcaveOn 𝕜 s (f • g) := by - rw [←neg_convexOn_iff, ←smul_neg] + rw [← neg_convexOn_iff, ← smul_neg] exact hf.smul' hg.neg hf₀ (fun x hx ↦ neg_nonneg.2 <| hg₀ hx) hfg.neg_right lemma ConcaveOn.smul_convexOn (hf : ConcaveOn 𝕜 s f) (hg : ConvexOn 𝕜 s g) (hf₀ : ∀ ⦃x⦄, x ∈ s → 0 ≤ f x) (hg₀ : ∀ ⦃x⦄, x ∈ s → g x ≤ 0) (hfg : MonovaryOn f g s) : ConvexOn 𝕜 s (f • g) := by - rw [←neg_concaveOn_iff, ←smul_neg] + rw [← neg_concaveOn_iff, ← smul_neg] exact hf.smul' hg.neg hf₀ (fun x hx ↦ neg_nonneg.2 <| hg₀ hx) hfg.neg_right lemma ConvexOn.smul_concaveOn' (hf : ConvexOn 𝕜 s f) (hg : ConcaveOn 𝕜 s g) (hf₀ : ∀ ⦃x⦄, x ∈ s → f x ≤ 0) (hg₀ : ∀ ⦃x⦄, x ∈ s → 0 ≤ g x) (hfg : MonovaryOn f g s) : ConvexOn 𝕜 s (f • g) := by - rw [←neg_concaveOn_iff, ←smul_neg] + rw [← neg_concaveOn_iff, ← smul_neg] exact hf.smul'' hg.neg hf₀ (fun x hx ↦ neg_nonpos.2 $ hg₀ hx) hfg.neg_right lemma ConcaveOn.smul_convexOn' (hf : ConcaveOn 𝕜 s f) (hg : ConvexOn 𝕜 s g) (hf₀ : ∀ ⦃x⦄, x ∈ s → f x ≤ 0) (hg₀ : ∀ ⦃x⦄, x ∈ s → 0 ≤ g x) (hfg : AntivaryOn f g s) : ConcaveOn 𝕜 s (f • g) := by - rw [←neg_convexOn_iff, ←smul_neg] + rw [← neg_convexOn_iff, ← smul_neg] exact hf.smul'' hg.neg hf₀ (fun x hx ↦ neg_nonpos.2 $ hg₀ hx) hfg.neg_right variable [IsScalarTower 𝕜 E E] [SMulCommClass 𝕜 E E] {f g : 𝕜 → E} @@ -146,7 +146,7 @@ lemma convexOn_pow : ∀ n, ConvexOn 𝕜 (Ici 0) fun x : 𝕜 ↦ x ^ n := /-- `x^n`, `n : ℕ` is convex on the whole real line whenever `n` is even. -/ protected lemma Even.convexOn_pow {n : ℕ} (hn : Even n) : ConvexOn 𝕜 univ fun x : 𝕜 ↦ x ^ n := by obtain ⟨n, rfl⟩ := hn - simp_rw [←two_mul, pow_mul] + simp_rw [← two_mul, pow_mul] refine ConvexOn.pow ⟨convex_univ, fun x _ y _ a b ha hb hab ↦ sub_nonneg.1 ?_⟩ (fun _ _ ↦ by positivity) _ calc @@ -165,11 +165,11 @@ lemma convexOn_zpow : ∀ n : ℤ, ConvexOn 𝕜 (Ioi 0) fun x : 𝕜 ↦ x ^ n simp_rw [zpow_ofNat] exact (convexOn_pow n).subset Ioi_subset_Ici_self (convex_Ioi _) | -[n+1] => by - simp_rw [zpow_negSucc, ←inv_pow] + simp_rw [zpow_negSucc, ← inv_pow] refine (convexOn_iff_forall_pos.2 ⟨convex_Ioi _, ?_⟩).pow (fun x (hx : 0 < x) ↦ by positivity) _ rintro x (hx : 0 < x) y (hy : 0 < y) a b ha hb hab field_simp - rw [div_le_div_iff, ←sub_nonneg] + rw [div_le_div_iff, ← sub_nonneg] calc 0 ≤ a * b * (x - y) ^ 2 := by positivity _ = _ := by obtain rfl := eq_sub_of_add_eq hab; ring diff --git a/Mathlib/Analysis/Convex/Radon.lean b/Mathlib/Analysis/Convex/Radon.lean index b26add9636f2c..274720066fb29 100644 --- a/Mathlib/Analysis/Convex/Radon.lean +++ b/Mathlib/Analysis/Convex/Radon.lean @@ -40,11 +40,11 @@ theorem radon_partition (h : ¬ AffineIndependent 𝕜 f) : exact sum_pos' (λ _i hi ↦ (mem_filter.1 hi).2) ⟨pos_w_index, by simp only [mem_filter, h1', h2'.le, and_self, h2']⟩ have hp : centerMass J w f = p := Finset.centerMass_of_sum_add_sum_eq_zero hJI $ by - simpa only [←h_vsum, not_lt] using sum_filter_add_sum_filter_not s (fun i ↦ w i < 0) _ + simpa only [← h_vsum, not_lt] using sum_filter_add_sum_filter_not s (fun i ↦ w i < 0) _ refine ⟨I, p, ?_, ?_⟩ · exact centerMass_mem_convexHull _ (fun _i hi ↦ (mem_filter.mp hi).2) hI (fun _i hi ↦ Set.mem_image_of_mem _ hi) - rw [←hp] + rw [← hp] refine centerMass_mem_convexHull_of_nonpos _ (fun _ hi ↦ (mem_filter.mp hi).2.le) ?_ (fun _i hi ↦ Set.mem_image_of_mem _ fun hi' ↦ ?_) · linarith only [hI, hJI] diff --git a/Mathlib/Analysis/Convex/SimplicialComplex/Basic.lean b/Mathlib/Analysis/Convex/SimplicialComplex/Basic.lean index fb1608f202286..8e9400a8bcbf4 100644 --- a/Mathlib/Analysis/Convex/SimplicialComplex/Basic.lean +++ b/Mathlib/Analysis/Convex/SimplicialComplex/Basic.lean @@ -107,7 +107,7 @@ theorem disjoint_or_exists_inter_eq_convexHull (hs : s ∈ K.faces) (ht : t ∈ Disjoint (convexHull 𝕜 (s : Set E)) (convexHull 𝕜 ↑t) ∨ ∃ u ∈ K.faces, convexHull 𝕜 (s : Set E) ∩ convexHull 𝕜 ↑t = convexHull 𝕜 ↑u := by classical - by_contra' h + by_contra! h refine' h.2 (s ∩ t) (K.down_closed hs (inter_subset_left _ _) fun hst => h.1 <| disjoint_iff_inf_le.mpr <| (K.inter_subset_convexHull hs ht).trans _) _ · rw [← coe_inter, hst, coe_empty, convexHull_empty] diff --git a/Mathlib/Analysis/Convex/SpecificFunctions/Pow.lean b/Mathlib/Analysis/Convex/SpecificFunctions/Pow.lean index 34a929cfe0b38..0f573fc5f8eff 100644 --- a/Mathlib/Analysis/Convex/SpecificFunctions/Pow.lean +++ b/Mathlib/Analysis/Convex/SpecificFunctions/Pow.lean @@ -42,7 +42,7 @@ lemma strictConcaveOn_rpow {p : ℝ} (hp₀ : 0 < p) (hp₁ : p < 1) : (by simp [hxy]) ha hb (by simp; norm_cast) have h₂ : ∀ x, f.symm x = x ^ p := by simp [NNReal.orderIsoRpow_symm_eq] refine ⟨convex_univ, fun x _ y _ hxy a b ha hb hab => ?_⟩ - simp only [←h₂] + simp only [← h₂] exact (f.strictConcaveOn_symm h₁).2 (Set.mem_univ x) (Set.mem_univ y) hxy ha hb hab lemma concaveOn_rpow {p : ℝ} (hp₀ : 0 ≤ p) (hp₁ : p ≤ 1) : diff --git a/Mathlib/Analysis/Convex/StoneSeparation.lean b/Mathlib/Analysis/Convex/StoneSeparation.lean index 0cd1081486832..7f238e873e683 100644 --- a/Mathlib/Analysis/Convex/StoneSeparation.lean +++ b/Mathlib/Analysis/Convex/StoneSeparation.lean @@ -101,7 +101,7 @@ theorem exists_convex_convex_compl_subset (hs : Convex 𝕜 s) (ht : Convex 𝕜 (hC.2.symm.mono (ht.segment_subset hut hvt) <| convexHull_min _ hC.1) simpa [insert_subset_iff, hp, hq, singleton_subset_iff.2 hzC] rintro c hc - by_contra' h + by_contra! h suffices h : Disjoint (convexHull 𝕜 (insert c C)) t · rw [← hCmax _ ⟨convex_convexHull _ _, h⟩ ((subset_insert _ _).trans <| subset_convexHull _ _)] at hc diff --git a/Mathlib/Analysis/Convex/Strict.lean b/Mathlib/Analysis/Convex/Strict.lean index e3cee8aebd9c1..3d44e3c9ad5a5 100644 --- a/Mathlib/Analysis/Convex/Strict.lean +++ b/Mathlib/Analysis/Convex/Strict.lean @@ -109,12 +109,13 @@ protected theorem StrictConvex.convex (hs : StrictConvex 𝕜 s) : Convex 𝕜 s #align strict_convex.convex StrictConvex.convex /-- An open convex set is strictly convex. -/ -protected theorem Convex.strictConvex_of_open (h : IsOpen s) (hs : Convex 𝕜 s) : StrictConvex 𝕜 s := +protected theorem Convex.strictConvex_of_isOpen (h : IsOpen s) (hs : Convex 𝕜 s) : + StrictConvex 𝕜 s := fun _ hx _ hy _ _ _ ha hb hab => h.interior_eq.symm ▸ hs hx hy ha.le hb.le hab -#align convex.strict_convex_of_open Convex.strictConvex_of_open +#align convex.strict_convex_of_open Convex.strictConvex_of_isOpen theorem IsOpen.strictConvex_iff (h : IsOpen s) : StrictConvex 𝕜 s ↔ Convex 𝕜 s := - ⟨StrictConvex.convex, Convex.strictConvex_of_open h⟩ + ⟨StrictConvex.convex, Convex.strictConvex_of_isOpen h⟩ #align is_open.strict_convex_iff IsOpen.strictConvex_iff theorem strictConvex_singleton (c : E) : StrictConvex 𝕜 ({c} : Set E) := diff --git a/Mathlib/Analysis/Convex/StrictConvexBetween.lean b/Mathlib/Analysis/Convex/StrictConvexBetween.lean index e4cf772256e97..ff588ee9b2890 100644 --- a/Mathlib/Analysis/Convex/StrictConvexBetween.lean +++ b/Mathlib/Analysis/Convex/StrictConvexBetween.lean @@ -111,8 +111,9 @@ variable {E F PE PF : Type*} [NormedAddCommGroup E] [NormedAddCommGroup F] [Norm lemma eq_lineMap_of_dist_eq_mul_of_dist_eq_mul (hxy : dist x y = r * dist x z) (hyz : dist y z = (1 - r) * dist x z) : y = AffineMap.lineMap x z r := by have : y -ᵥ x ∈ [(0 : E) -[ℝ] z -ᵥ x] := by - rw [mem_segment_iff_wbtw, ←dist_add_dist_eq_iff, dist_zero_left, dist_vsub_cancel_right, - ←dist_eq_norm_vsub', ←dist_eq_norm_vsub', hxy, hyz, ←add_mul, add_sub_cancel'_right, one_mul] + rw [mem_segment_iff_wbtw, ← dist_add_dist_eq_iff, dist_zero_left, dist_vsub_cancel_right, + ← dist_eq_norm_vsub', ← dist_eq_norm_vsub', hxy, hyz, ← add_mul, add_sub_cancel'_right, + one_mul] obtain rfl | hne := eq_or_ne x z · obtain rfl : y = x := by simpa simp diff --git a/Mathlib/Analysis/Convex/Strong.lean b/Mathlib/Analysis/Convex/Strong.lean index 19a08901c13cc..dd0aa2e5f2834 100644 --- a/Mathlib/Analysis/Convex/Strong.lean +++ b/Mathlib/Analysis/Convex/Strong.lean @@ -72,7 +72,7 @@ lemma UniformConvexOn.strictConvexOn (hf : UniformConvexOn s φ f) (hφ : ∀ r, StrictConvexOn ℝ s f := by refine ⟨hf.1, fun x hx y hy hxy a b ha hb hab ↦ (hf.2 hx hy ha.le hb.le hab).trans_lt $ sub_lt_self _ ?_⟩ - rw [←sub_ne_zero, ←norm_pos_iff] at hxy + rw [← sub_ne_zero, ← norm_pos_iff] at hxy have := hφ _ hxy.ne' positivity @@ -80,7 +80,7 @@ lemma UniformConcaveOn.strictConcaveOn (hf : UniformConcaveOn s φ f) (hφ : ∀ StrictConcaveOn ℝ s f := by refine ⟨hf.1, fun x hx y hy hxy a b ha hb hab ↦ (hf.2 hx hy ha.le hb.le hab).trans_lt' $ lt_add_of_pos_right _ ?_⟩ - rw [←sub_ne_zero, ←norm_pos_iff] at hxy + rw [← sub_ne_zero, ← norm_pos_iff] at hxy have := hφ _ hxy.ne' positivity @@ -101,7 +101,7 @@ lemma UniformConvexOn.neg (hf : UniformConvexOn s φ f) : UniformConcaveOn s φ lemma UniformConcaveOn.neg (hf : UniformConcaveOn s φ f) : UniformConvexOn s φ (-f) := by refine ⟨hf.1, fun x hx y hy a b ha hb hab ↦ le_of_neg_le_neg ?_⟩ - simpa [add_comm, -neg_le_neg_iff, ←le_sub_iff_add_le', sub_eq_add_neg, neg_add] + simpa [add_comm, -neg_le_neg_iff, ← le_sub_iff_add_le', sub_eq_add_neg, neg_add] using hf.2 hx hy ha hb hab lemma UniformConvexOn.sub (hf : UniformConvexOn s φ f) (hg : UniformConcaveOn s ψ g) : @@ -134,10 +134,10 @@ nonrec lemma StrongConcaveOn.mono (hmn : m ≤ n) (hf : StrongConcaveOn s n f) : hf.mono fun r ↦ mul_le_mul_of_nonneg_right (div_le_div_of_le zero_le_two hmn) $ by positivity @[simp] lemma strongConvexOn_zero : StrongConvexOn s 0 f ↔ ConvexOn ℝ s f := by - simp [StrongConvexOn, ←Pi.zero_def] + simp [StrongConvexOn, ← Pi.zero_def] @[simp] lemma strongConcaveOn_zero : StrongConcaveOn s 0 f ↔ ConcaveOn ℝ s f := by - simp [StrongConcaveOn, ←Pi.zero_def] + simp [StrongConcaveOn, ← Pi.zero_def] nonrec lemma StrongConvexOn.strictConvexOn (hf : StrongConvexOn s m f) (hm : 0 < m) : StrictConvexOn ℝ s f := hf.strictConvexOn fun r hr ↦ by positivity @@ -173,6 +173,6 @@ lemma strongConvexOn_iff_convex : lemma strongConcaveOn_iff_convex : StrongConcaveOn s m f ↔ ConcaveOn ℝ s fun x ↦ f x + m / (2 : ℝ) * ‖x‖ ^ 2 := by refine and_congr_right fun _ ↦ forall₄_congr fun x _ y _ ↦ forall₅_congr fun a b ha hb hab ↦ ?_ - simp_rw [←sub_le_iff_le_add, smul_eq_mul, aux_add ha hb hab, mul_assoc, mul_left_comm] + simp_rw [← sub_le_iff_le_add, smul_eq_mul, aux_add ha hb hab, mul_assoc, mul_left_comm] end InnerProductSpace diff --git a/Mathlib/Analysis/Convex/Topology.lean b/Mathlib/Analysis/Convex/Topology.lean index 6be7c69092547..77b954c50b5ea 100644 --- a/Mathlib/Analysis/Convex/Topology.lean +++ b/Mathlib/Analysis/Convex/Topology.lean @@ -333,7 +333,7 @@ theorem Convex.subset_interior_image_homothety_of_one_lt {s : Set E} (hs : Conve subset_closure.trans <| hs.closure_subset_interior_image_homothety_of_one_lt hx t ht #align convex.subset_interior_image_homothety_of_one_lt Convex.subset_interior_image_homothety_of_one_lt -theorem JoinedIn_of_segment_subset {E : Type*} [AddCommGroup E] [Module ℝ E] +theorem JoinedIn.of_segment_subset {E : Type*} [AddCommGroup E] [Module ℝ E] [TopologicalSpace E] [ContinuousAdd E] [ContinuousSMul ℝ E] {x y : E} {s : Set E} (h : [x -[ℝ] y] ⊆ s) : JoinedIn s x y := by have A : Continuous (fun t ↦ (1 - t) • x + t • y : ℝ → E) := by continuity @@ -346,7 +346,7 @@ protected theorem Convex.isPathConnected {s : Set E} (hconv : Convex ℝ s) (hne IsPathConnected s := by refine' isPathConnected_iff.mpr ⟨hne, _⟩ intro x x_in y y_in - exact JoinedIn_of_segment_subset ((segment_subset_iff ℝ).2 (hconv x_in y_in)) + exact JoinedIn.of_segment_subset ((segment_subset_iff ℝ).2 (hconv x_in y_in)) #align convex.is_path_connected Convex.isPathConnected /-- A nonempty convex set is connected. -/ diff --git a/Mathlib/Analysis/Convolution.lean b/Mathlib/Analysis/Convolution.lean index 698c702afe91b..40a9654ab48bf 100644 --- a/Mathlib/Analysis/Convolution.lean +++ b/Mathlib/Analysis/Convolution.lean @@ -1172,7 +1172,7 @@ theorem hasFDerivAt_convolution_right_with_param {g : P → G → E'} {s : Set P have A : IsCompact ({q₀.1} ×ˢ k) := isCompact_singleton.prod hk obtain ⟨t, kt, t_open, ht⟩ : ∃ t, {q₀.1} ×ˢ k ⊆ t ∧ IsOpen t ∧ IsBounded (g' '' t) := by have B : ContinuousOn g' (s ×ˢ univ) := - hg.continuousOn_fderiv_of_open (hs.prod isOpen_univ) le_rfl + hg.continuousOn_fderiv_of_isOpen (hs.prod isOpen_univ) le_rfl apply exists_isOpen_isBounded_image_of_isCompact_of_continuousOn A (hs.prod isOpen_univ) _ B simp only [prod_subset_prod_iff, hq₀, singleton_subset_iff, subset_univ, and_self_iff, true_or_iff] @@ -1220,7 +1220,7 @@ theorem hasFDerivAt_convolution_right_with_param {g : P → G → E'} {s : Set P HasCompactSupport.intro hk fun x hx => g'_zero q₀.1 x hq₀ hx apply (HasCompactSupport.convolutionExists_right (L.precompR (P × G) : _) T hf _ q₀.2).1 have : ContinuousOn g' (s ×ˢ univ) := - hg.continuousOn_fderiv_of_open (hs.prod isOpen_univ) le_rfl + hg.continuousOn_fderiv_of_isOpen (hs.prod isOpen_univ) le_rfl apply this.comp_continuous (continuous_const.prod_mk continuous_id') intro x simpa only [prod_mk_mem_set_prod_eq, mem_univ, and_true_iff] using hq₀ @@ -1310,7 +1310,7 @@ theorem contDiffOn_convolution_right_with_param_aux {G : Type uP} {E' : Type uP} have A : ∀ q₀ : P × G, q₀.1 ∈ s → HasFDerivAt (fun q : P × G => (f ⋆[L, μ] g q.1) q.2) (f' q₀.1 q₀.2) q₀ := hasFDerivAt_convolution_right_with_param L hs hk hgs hf hg.one_of_succ - rw [contDiffOn_succ_iff_fderiv_of_open (hs.prod (@isOpen_univ G _))] at hg ⊢ + rw [contDiffOn_succ_iff_fderiv_of_isOpen (hs.prod (@isOpen_univ G _))] at hg ⊢ constructor · rintro ⟨p, x⟩ ⟨hp, -⟩ exact (A (p, x) hp).differentiableAt.differentiableWithinAt diff --git a/Mathlib/Analysis/InnerProductSpace/Basic.lean b/Mathlib/Analysis/InnerProductSpace/Basic.lean index bee26b60e69f8..7a9c7c1d2efc5 100644 --- a/Mathlib/Analysis/InnerProductSpace/Basic.lean +++ b/Mathlib/Analysis/InnerProductSpace/Basic.lean @@ -1892,7 +1892,7 @@ theorem Orthonormal.sum_inner_products_le {s : Finset ι} (hv : Orthonormal 𝕜 rw [@norm_sub_sq 𝕜, sub_add] classical simp only [@InnerProductSpace.norm_sq_eq_inner 𝕜, _root_.inner_sum, _root_.sum_inner] - simp only [inner_smul_right, two_mul, inner_smul_left, inner_conj_symm, ←mul_assoc, h₂, + simp only [inner_smul_right, two_mul, inner_smul_left, inner_conj_symm, ← mul_assoc, h₂, add_sub_cancel, sub_right_inj] simp only [map_sum, ← inner_conj_symm x, ← h₃] diff --git a/Mathlib/Analysis/InnerProductSpace/Spectrum.lean b/Mathlib/Analysis/InnerProductSpace/Spectrum.lean index 0767d63273edf..58de7527eb1cb 100644 --- a/Mathlib/Analysis/InnerProductSpace/Spectrum.lean +++ b/Mathlib/Analysis/InnerProductSpace/Spectrum.lean @@ -289,7 +289,7 @@ theorem eigenvalue_nonneg_of_nonneg {μ : ℝ} {T : E →ₗ[𝕜] E} (hμ : Has have : IsROrC.re ⟪v, T v⟫ = μ * ‖v‖ ^ 2 := by have := congr_arg IsROrC.re (inner_product_apply_eigenvector hv.1) -- porting note: why can't `exact_mod_cast` do this? These lemmas are marked `norm_cast` - rw [←IsROrC.ofReal_pow, ←IsROrC.ofReal_mul] at this + rw [← IsROrC.ofReal_pow, ← IsROrC.ofReal_mul] at this exact mod_cast this exact (zero_le_mul_right hpos).mp (this ▸ hnn v) #align eigenvalue_nonneg_of_nonneg eigenvalue_nonneg_of_nonneg @@ -301,7 +301,7 @@ theorem eigenvalue_pos_of_pos {μ : ℝ} {T : E →ₗ[𝕜] E} (hμ : HasEigenv have : IsROrC.re ⟪v, T v⟫ = μ * ‖v‖ ^ 2 := by have := congr_arg IsROrC.re (inner_product_apply_eigenvector hv.1) -- porting note: why can't `exact_mod_cast` do this? These lemmas are marked `norm_cast` - rw [←IsROrC.ofReal_pow, ←IsROrC.ofReal_mul] at this + rw [← IsROrC.ofReal_pow, ← IsROrC.ofReal_mul] at this exact mod_cast this exact (zero_lt_mul_right hpos).mp (this ▸ hnn v) #align eigenvalue_pos_of_pos eigenvalue_pos_of_pos diff --git a/Mathlib/Analysis/LocallyConvex/AbsConvex.lean b/Mathlib/Analysis/LocallyConvex/AbsConvex.lean index bce4e0375ac1f..2bdde66f88352 100644 --- a/Mathlib/Analysis/LocallyConvex/AbsConvex.lean +++ b/Mathlib/Analysis/LocallyConvex/AbsConvex.lean @@ -152,7 +152,7 @@ theorem gaugeSeminormFamily_ball (s : AbsConvexOpenSets 𝕜 E) : dsimp only [gaugeSeminormFamily] rw [Seminorm.ball_zero_eq] simp_rw [gaugeSeminorm_toFun] - exact gauge_lt_one_eq_self_of_open s.coe_convex s.coe_zero_mem s.coe_isOpen + exact gauge_lt_one_eq_self_of_isOpen s.coe_convex s.coe_zero_mem s.coe_isOpen #align gauge_seminorm_family_ball gaugeSeminormFamily_ball variable [TopologicalAddGroup E] [ContinuousSMul 𝕜 E] diff --git a/Mathlib/Analysis/LocallyConvex/Bounded.lean b/Mathlib/Analysis/LocallyConvex/Bounded.lean index 03a794f0752c9..54236bc00f595 100644 --- a/Mathlib/Analysis/LocallyConvex/Bounded.lean +++ b/Mathlib/Analysis/LocallyConvex/Bounded.lean @@ -162,7 +162,7 @@ theorem isVonNBounded_of_smul_tendsto_zero {ε : ι → 𝕝} {l : Filter ι} [l (hε : ∀ᶠ n in l, ε n ≠ 0) {S : Set E} (H : ∀ x : ι → E, (∀ n, x n ∈ S) → Tendsto (ε • x) l (𝓝 0)) : IsVonNBounded 𝕝 S := by rw [(nhds_basis_balanced 𝕝 E).isVonNBounded_basis_iff] - by_contra' H' + by_contra! H' rcases H' with ⟨V, ⟨hV, hVb⟩, hVS⟩ have : ∀ᶠ n in l, ∃ x : S, ε n • (x : E) ∉ V := by filter_upwards [hε] with n hn diff --git a/Mathlib/Analysis/LocallyConvex/WithSeminorms.lean b/Mathlib/Analysis/LocallyConvex/WithSeminorms.lean index bb842482e17cb..103535f3ba76b 100644 --- a/Mathlib/Analysis/LocallyConvex/WithSeminorms.lean +++ b/Mathlib/Analysis/LocallyConvex/WithSeminorms.lean @@ -354,7 +354,7 @@ theorem WithSeminorms.T1_of_separating (hp : WithSeminorms p) theorem WithSeminorms.separating_of_T1 [T1Space E] (hp : WithSeminorms p) (x : E) (hx : x ≠ 0) : ∃ i, p i x ≠ 0 := by have := ((t1Space_TFAE E).out 0 9).mp (inferInstanceAs <| T1Space E) - by_contra' h + by_contra! h refine' hx (this _) rw [hp.hasBasis_zero_ball.specializes_iff] rintro ⟨s, r⟩ (hr : 0 < r) diff --git a/Mathlib/Analysis/Normed/Field/Basic.lean b/Mathlib/Analysis/Normed/Field/Basic.lean index 288d23789de85..426ecac3abf1a 100644 --- a/Mathlib/Analysis/Normed/Field/Basic.lean +++ b/Mathlib/Analysis/Normed/Field/Basic.lean @@ -765,12 +765,12 @@ instance (priority := 100) NormedField.toNormedCommRing : NormedCommRing α := @[simp] theorem norm_prod (s : Finset β) (f : β → α) : ‖∏ b in s, f b‖ = ∏ b in s, ‖f b‖ := - (normHom.toMonoidHom : α →* ℝ).map_prod f s + map_prod normHom.toMonoidHom f s #align norm_prod norm_prod @[simp] theorem nnnorm_prod (s : Finset β) (f : β → α) : ‖∏ b in s, f b‖₊ = ∏ b in s, ‖f b‖₊ := - (nnnormHom.toMonoidHom : α →* ℝ≥0).map_prod f s + map_prod nnnormHom.toMonoidHom f s #align nnnorm_prod nnnorm_prod end NormedField @@ -949,7 +949,7 @@ instance Rat.normedField : NormedField ℚ := instance Rat.denselyNormedField : DenselyNormedField ℚ where lt_norm_lt r₁ r₂ h₀ hr := let ⟨q, h⟩ := exists_rat_btwn hr - ⟨q, by rwa [←Rat.norm_cast_real, Real.norm_eq_abs, abs_of_pos (h₀.trans_lt h.1)]⟩ + ⟨q, by rwa [← Rat.norm_cast_real, Real.norm_eq_abs, abs_of_pos (h₀.trans_lt h.1)]⟩ section RingHomIsometric variable {R₁ : Type*} {R₂ : Type*} {R₃ : Type*} diff --git a/Mathlib/Analysis/Normed/Group/AddTorsor.lean b/Mathlib/Analysis/Normed/Group/AddTorsor.lean index 326801912f1be..5169edbdbe15c 100644 --- a/Mathlib/Analysis/Normed/Group/AddTorsor.lean +++ b/Mathlib/Analysis/Normed/Group/AddTorsor.lean @@ -207,7 +207,7 @@ def pseudoMetricSpaceOfNormedAddCommGroupOfAddTorsor (V P : Type*) [SeminormedAd [AddTorsor V P] : PseudoMetricSpace P where dist x y := ‖(x -ᵥ y : V)‖ -- porting note: `edist_dist` is no longer an `autoParam` - edist_dist _ _ := by simp only [←ENNReal.ofReal_eq_coe_nnreal] + edist_dist _ _ := by simp only [← ENNReal.ofReal_eq_coe_nnreal] dist_self x := by simp dist_comm x y := by simp only [← neg_vsub_eq_vsub_rev y x, norm_neg] dist_triangle x y z := by diff --git a/Mathlib/Analysis/Normed/Group/Basic.lean b/Mathlib/Analysis/Normed/Group/Basic.lean index d09980f0cabb7..7a0b704151926 100644 --- a/Mathlib/Analysis/Normed/Group/Basic.lean +++ b/Mathlib/Analysis/Normed/Group/Basic.lean @@ -182,15 +182,16 @@ instance as a special case of a more general `SeminormedGroup` instance. -/ satisfying `∀ x, ‖x‖ = 0 → x = 0`. This avoids having to go back to the `(Pseudo)MetricSpace` level when declaring a `NormedAddGroup` instance as a special case of a more general `SeminormedAddGroup` instance."] -def NormedGroup.ofSeparation [SeminormedGroup E] (h : ∀ x : E, ‖x‖ = 0 → x = 1) : NormedGroup E := - { ‹SeminormedGroup E› with - toMetricSpace := - { eq_of_dist_eq_zero := fun hxy => - div_eq_one.1 <| h _ <| by exact (‹SeminormedGroup E›.dist_eq _ _).symm.trans hxy } } - -- porting note: the `rwa` no longer worked, but it was easy enough to provide the term. - -- however, notice that if you make `x` and `y` accessible, then the following does work: - -- `have := ‹SeminormedGroup E›.dist_eq x y; rwa [←this]`, so I'm not sure why the `rwa` - -- was broken. +def NormedGroup.ofSeparation [SeminormedGroup E] (h : ∀ x : E, ‖x‖ = 0 → x = 1) : + NormedGroup E where + dist_eq := ‹SeminormedGroup E›.dist_eq + toMetricSpace := + { eq_of_dist_eq_zero := fun hxy => + div_eq_one.1 <| h _ <| by exact (‹SeminormedGroup E›.dist_eq _ _).symm.trans hxy } + -- porting note: the `rwa` no longer worked, but it was easy enough to provide the term. + -- however, notice that if you make `x` and `y` accessible, then the following does work: + -- `have := ‹SeminormedGroup E›.dist_eq x y; rwa [← this]`, so I'm not sure why the `rwa` + -- was broken. #align normed_group.of_separation NormedGroup.ofSeparation #align normed_add_group.of_separation NormedAddGroup.ofSeparation diff --git a/Mathlib/Analysis/Normed/Group/Hom.lean b/Mathlib/Analysis/Normed/Group/Hom.lean index 462b117c6022d..cea2d02f62e43 100644 --- a/Mathlib/Analysis/Normed/Group/Hom.lean +++ b/Mathlib/Analysis/Normed/Group/Hom.lean @@ -153,7 +153,7 @@ theorem coe_toAddMonoidHom : ⇑f.toAddMonoidHom = f := theorem toAddMonoidHom_injective : Function.Injective (@NormedAddGroupHom.toAddMonoidHom V₁ V₂ _ _) := fun f g h => - coe_inj <| by rw [←coe_toAddMonoidHom f, ←coe_toAddMonoidHom g, h] + coe_inj <| by rw [← coe_toAddMonoidHom f, ← coe_toAddMonoidHom g, h] #align normed_add_group_hom.to_add_monoid_hom_injective NormedAddGroupHom.toAddMonoidHom_injective @[simp] @@ -619,7 +619,7 @@ def coeAddHom : NormedAddGroupHom V₁ V₂ →+ V₁ → V₂ where @[simp] theorem coe_sum {ι : Type*} (s : Finset ι) (f : ι → NormedAddGroupHom V₁ V₂) : ⇑(∑ i in s, f i) = ∑ i in s, (f i : V₁ → V₂) := - (coeAddHom : _ →+ V₁ → V₂).map_sum f s + map_sum coeAddHom f s #align normed_add_group_hom.coe_sum NormedAddGroupHom.coe_sum theorem sum_apply {ι : Type*} (s : Finset ι) (f : ι → NormedAddGroupHom V₁ V₂) (v : V₁) : @@ -747,7 +747,7 @@ theorem mem_ker (v : V₁) : v ∈ f.ker ↔ f v = 0 := by the corestriction of `f` to the kernel of `g`. -/ @[simps] def ker.lift (h : g.comp f = 0) : NormedAddGroupHom V₁ g.ker where - toFun v := ⟨f v, by rw [g.mem_ker, ←comp_apply g f, h, zero_apply]⟩ + toFun v := ⟨f v, by rw [g.mem_ker, ← comp_apply g f, h, zero_apply]⟩ map_add' v w := by simp only [map_add, AddSubmonoid.mk_add_mk] bound' := f.bound' #align normed_add_group_hom.ker.lift NormedAddGroupHom.ker.lift diff --git a/Mathlib/Analysis/Normed/Group/SemiNormedGroupCat/Kernels.lean b/Mathlib/Analysis/Normed/Group/SemiNormedGroupCat/Kernels.lean index ef56e75c0d415..ed92402c67426 100644 --- a/Mathlib/Analysis/Normed/Group/SemiNormedGroupCat/Kernels.lean +++ b/Mathlib/Analysis/Normed/Group/SemiNormedGroupCat/Kernels.lean @@ -162,7 +162,7 @@ def cokernelCocone {X Y : SemiNormedGroupCat.{u}} (f : X ⟶ Y) : Cofork f 0 := (inferInstance : SeminormedAddCommGroup Y) -- porting note: again simp doesn't seem to be firing in the below line -- This used to be `rw`, but we need `erw` after leanprover/lean4#2644 - erw [ ←NormedAddGroupHom.mem_ker, f.range.ker_normedMk, f.mem_range] + erw [← NormedAddGroupHom.mem_ker, f.range.ker_normedMk, f.mem_range] -- This used to be `simp only [exists_apply_eq_apply]` before leanprover/lean4#2644 convert exists_apply_eq_apply f a) set_option linter.uppercaseLean3 false in diff --git a/Mathlib/Analysis/Normed/Group/Seminorm.lean b/Mathlib/Analysis/Normed/Group/Seminorm.lean index eae564184226e..452e183e38454 100644 --- a/Mathlib/Analysis/Normed/Group/Seminorm.lean +++ b/Mathlib/Analysis/Normed/Group/Seminorm.lean @@ -643,7 +643,7 @@ instance : SMul R (GroupSeminorm E) := simp only [← smul_one_smul ℝ≥0 r (_ : ℝ), NNReal.smul_def, smul_eq_mul, map_one_eq_zero p, mul_zero] mul_le' := fun _ _ => by - simp only [← smul_one_smul ℝ≥0 r (_ : ℝ), NNReal.smul_def, smul_eq_mul, ←mul_add] + simp only [← smul_one_smul ℝ≥0 r (_ : ℝ), NNReal.smul_def, smul_eq_mul, ← mul_add] gcongr apply map_mul_le_add inv' := fun x => by simp_rw [map_inv_eq_map p] }⟩ diff --git a/Mathlib/Analysis/NormedSpace/Basic.lean b/Mathlib/Analysis/NormedSpace/Basic.lean index 6eea6485dc2c8..43b0c9ec47fd7 100644 --- a/Mathlib/Analysis/NormedSpace/Basic.lean +++ b/Mathlib/Analysis/NormedSpace/Basic.lean @@ -170,7 +170,7 @@ instance NormedSpace.discreteTopology_zmultiples rcases eq_or_ne e 0 with (rfl | he) · rw [AddSubgroup.zmultiples_zero_eq_bot] refine Subsingleton.discreteTopology (α := ↑(⊥ : Subspace ℚ E)) - · rw [discreteTopology_iff_open_singleton_zero, isOpen_induced_iff] + · rw [discreteTopology_iff_isOpen_singleton_zero, isOpen_induced_iff] refine' ⟨Metric.ball 0 ‖e‖, Metric.isOpen_ball, _⟩ ext ⟨x, hx⟩ obtain ⟨k, rfl⟩ := AddSubgroup.mem_zmultiples_iff.mp hx diff --git a/Mathlib/Analysis/NormedSpace/BoundedLinearMaps.lean b/Mathlib/Analysis/NormedSpace/BoundedLinearMaps.lean index a8811e967d331..c6b19721488cf 100644 --- a/Mathlib/Analysis/NormedSpace/BoundedLinearMaps.lean +++ b/Mathlib/Analysis/NormedSpace/BoundedLinearMaps.lean @@ -3,7 +3,7 @@ Copyright (c) 2018 Patrick Massot. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Patrick Massot, Johannes Hölzl -/ -import Mathlib.Analysis.NormedSpace.Multilinear +import Mathlib.Analysis.NormedSpace.Multilinear.Basic import Mathlib.Analysis.NormedSpace.Units import Mathlib.Analysis.Asymptotics.Asymptotics @@ -39,9 +39,10 @@ is normed) that `‖f x‖` is bounded by a multiple of `‖x‖`. Hence the "bo ## Notes -The main use of this file is `IsBoundedBilinearMap`. The file `Analysis.NormedSpace.Multilinear` -already expounds the theory of multilinear maps, but the `2`-variables case is sufficiently simpler -to currently deserve its own treatment. +The main use of this file is `IsBoundedBilinearMap`. +The file `Analysis.NormedSpace.Multilinear.Basic` +already expounds the theory of multilinear maps, +but the `2`-variables case is sufficiently simpler to currently deserve its own treatment. `IsBoundedLinearMap` is effectively an unbundled version of `ContinuousLinearMap` (defined in `Topology.Algebra.Module.Basic`, theory over normed spaces developed in diff --git a/Mathlib/Analysis/NormedSpace/CompactOperator.lean b/Mathlib/Analysis/NormedSpace/CompactOperator.lean index 0d210e46a9fc1..97e90601a36ce 100644 --- a/Mathlib/Analysis/NormedSpace/CompactOperator.lean +++ b/Mathlib/Analysis/NormedSpace/CompactOperator.lean @@ -434,8 +434,7 @@ theorem isClosed_setOf_isCompactOperator {𝕜₁ 𝕜₂ : Type*} [Nontrivially theorem compactOperator_topologicalClosure {𝕜₁ 𝕜₂ : Type*} [NontriviallyNormedField 𝕜₁] [NormedField 𝕜₂] {σ₁₂ : 𝕜₁ →+* 𝕜₂} {M₁ M₂ : Type*} [SeminormedAddCommGroup M₁] [AddCommGroup M₂] [NormedSpace 𝕜₁ M₁] [Module 𝕜₂ M₂] [UniformSpace M₂] [UniformAddGroup M₂] - [ContinuousConstSMul 𝕜₂ M₂] [T2Space M₂] [CompleteSpace M₂] - [ContinuousSMul 𝕜₂ (M₁ →SL[σ₁₂] M₂)] : + [ContinuousConstSMul 𝕜₂ M₂] [T2Space M₂] [CompleteSpace M₂] : (compactOperator σ₁₂ M₁ M₂).topologicalClosure = compactOperator σ₁₂ M₁ M₂ := SetLike.ext' isClosed_setOf_isCompactOperator.closure_eq #align compact_operator_topological_closure compactOperator_topologicalClosure diff --git a/Mathlib/Analysis/NormedSpace/Complemented.lean b/Mathlib/Analysis/NormedSpace/Complemented.lean index 932044711800e..62efcfe659b47 100644 --- a/Mathlib/Analysis/NormedSpace/Complemented.lean +++ b/Mathlib/Analysis/NormedSpace/Complemented.lean @@ -77,7 +77,7 @@ theorem equivProdOfSurjectiveOfIsCompl_apply {f : E →L[𝕜] F} {g : E →L[ end ContinuousLinearMap -namespace Subspace +namespace Submodule variable [CompleteSpace E] (p q : Subspace 𝕜 E) @@ -88,13 +88,13 @@ def prodEquivOfClosedCompl (h : IsCompl p q) (hp : IsClosed (p : Set E)) haveI := hp.completeSpace_coe; haveI := hq.completeSpace_coe refine' (p.prodEquivOfIsCompl q h).toContinuousLinearEquivOfContinuous _ exact (p.subtypeL.coprod q.subtypeL).continuous -#align subspace.prod_equiv_of_closed_compl Subspace.prodEquivOfClosedCompl +#align subspace.prod_equiv_of_closed_compl Submodule.prodEquivOfClosedCompl /-- Projection to a closed submodule along a closed complement. -/ def linearProjOfClosedCompl (h : IsCompl p q) (hp : IsClosed (p : Set E)) (hq : IsClosed (q : Set E)) : E →L[𝕜] p := ContinuousLinearMap.fst 𝕜 p q ∘L ↑(prodEquivOfClosedCompl p q h hp hq).symm -#align subspace.linear_proj_of_closed_compl Subspace.linearProjOfClosedCompl +#align subspace.linear_proj_of_closed_compl Submodule.linearProjOfClosedCompl variable {p q} @@ -102,43 +102,45 @@ variable {p q} theorem coe_prodEquivOfClosedCompl (h : IsCompl p q) (hp : IsClosed (p : Set E)) (hq : IsClosed (q : Set E)) : ⇑(p.prodEquivOfClosedCompl q h hp hq) = p.prodEquivOfIsCompl q h := rfl -#align subspace.coe_prod_equiv_of_closed_compl Subspace.coe_prodEquivOfClosedCompl +#align subspace.coe_prod_equiv_of_closed_compl Submodule.coe_prodEquivOfClosedCompl @[simp] theorem coe_prodEquivOfClosedCompl_symm (h : IsCompl p q) (hp : IsClosed (p : Set E)) (hq : IsClosed (q : Set E)) : ⇑(p.prodEquivOfClosedCompl q h hp hq).symm = (p.prodEquivOfIsCompl q h).symm := rfl -#align subspace.coe_prod_equiv_of_closed_compl_symm Subspace.coe_prodEquivOfClosedCompl_symm +#align subspace.coe_prod_equiv_of_closed_compl_symm Submodule.coe_prodEquivOfClosedCompl_symm @[simp] theorem coe_continuous_linearProjOfClosedCompl (h : IsCompl p q) (hp : IsClosed (p : Set E)) (hq : IsClosed (q : Set E)) : (p.linearProjOfClosedCompl q h hp hq : E →ₗ[𝕜] p) = p.linearProjOfIsCompl q h := rfl -#align subspace.coe_continuous_linear_proj_of_closed_compl Subspace.coe_continuous_linearProjOfClosedCompl +#align subspace.coe_continuous_linear_proj_of_closed_compl Submodule.coe_continuous_linearProjOfClosedCompl @[simp] theorem coe_continuous_linearProjOfClosedCompl' (h : IsCompl p q) (hp : IsClosed (p : Set E)) (hq : IsClosed (q : Set E)) : ⇑(p.linearProjOfClosedCompl q h hp hq) = p.linearProjOfIsCompl q h := rfl -#align subspace.coe_continuous_linear_proj_of_closed_compl' Subspace.coe_continuous_linearProjOfClosedCompl' +#align subspace.coe_continuous_linear_proj_of_closed_compl' Submodule.coe_continuous_linearProjOfClosedCompl' -theorem closedComplemented_of_closed_compl (h : IsCompl p q) (hp : IsClosed (p : Set E)) +theorem ClosedComplemented.of_isCompl_isClosed (h : IsCompl p q) (hp : IsClosed (p : Set E)) (hq : IsClosed (q : Set E)) : p.ClosedComplemented := ⟨p.linearProjOfClosedCompl q h hp hq, Submodule.linearProjOfIsCompl_apply_left h⟩ -#align subspace.closed_complemented_of_closed_compl Subspace.closedComplemented_of_closed_compl +#align subspace.closed_complemented_of_closed_compl Submodule.ClosedComplemented.of_isCompl_isClosed -theorem closedComplemented_iff_has_closed_compl : +alias IsCompl.closedComplemented_of_isClosed := ClosedComplemented.of_isCompl_isClosed + +theorem closedComplemented_iff_isClosed_exists_isClosed_isCompl : p.ClosedComplemented ↔ - IsClosed (p : Set E) ∧ ∃ (q : Subspace 𝕜 E) (_ : IsClosed (q : Set E)), IsCompl p q := - ⟨fun h => ⟨h.isClosed, h.has_closed_complement⟩, - fun ⟨hp, ⟨_, hq, hpq⟩⟩ => closedComplemented_of_closed_compl hpq hp hq⟩ -#align subspace.closed_complemented_iff_has_closed_compl Subspace.closedComplemented_iff_has_closed_compl + IsClosed (p : Set E) ∧ ∃ q : Submodule 𝕜 E, IsClosed (q : Set E) ∧ IsCompl p q := + ⟨fun h => ⟨h.isClosed, h.exists_isClosed_isCompl⟩, + fun ⟨hp, ⟨_, hq, hpq⟩⟩ => .of_isCompl_isClosed hpq hp hq⟩ +#align subspace.closed_complemented_iff_has_closed_compl Submodule.closedComplemented_iff_isClosed_exists_isClosed_isCompl -theorem closedComplemented_of_quotient_finiteDimensional [CompleteSpace 𝕜] +theorem ClosedComplemented.of_quotient_finiteDimensional [CompleteSpace 𝕜] [FiniteDimensional 𝕜 (E ⧸ p)] (hp : IsClosed (p : Set E)) : p.ClosedComplemented := by obtain ⟨q, hq⟩ : ∃ q, IsCompl p q := p.exists_isCompl haveI : FiniteDimensional 𝕜 q := (p.quotientEquivOfIsCompl q hq).finiteDimensional - exact closedComplemented_of_closed_compl hq hp q.closed_of_finiteDimensional -#align subspace.closed_complemented_of_quotient_finite_dimensional Subspace.closedComplemented_of_quotient_finiteDimensional + exact .of_isCompl_isClosed hq hp q.closed_of_finiteDimensional +#align subspace.closed_complemented_of_quotient_finite_dimensional Submodule.ClosedComplemented.of_quotient_finiteDimensional -end Subspace +end Submodule diff --git a/Mathlib/Analysis/NormedSpace/Connected.lean b/Mathlib/Analysis/NormedSpace/Connected.lean index 25b48e272af4e..f3281222be9d1 100644 --- a/Mathlib/Analysis/NormedSpace/Connected.lean +++ b/Mathlib/Analysis/NormedSpace/Connected.lean @@ -89,12 +89,12 @@ theorem Set.Countable.isPathConnected_compl_of_one_lt_rank simp only [compl_union, mem_inter_iff, mem_compl_iff, mem_setOf_eq, not_nonempty_iff_eq_empty] at ht have JA : JoinedIn sᶜ a z := by - apply JoinedIn_of_segment_subset + apply JoinedIn.of_segment_subset rw [subset_compl_iff_disjoint_right, disjoint_iff_inter_eq_empty] convert ht.2 exact Ia.symm have JB : JoinedIn sᶜ b z := by - apply JoinedIn_of_segment_subset + apply JoinedIn.of_segment_subset rw [subset_compl_iff_disjoint_right, disjoint_iff_inter_eq_empty] convert ht.1 exact Ib.symm diff --git a/Mathlib/Analysis/NormedSpace/HahnBanach/Extension.lean b/Mathlib/Analysis/NormedSpace/HahnBanach/Extension.lean index a7343360748bf..6dfa6ca5bab72 100644 --- a/Mathlib/Analysis/NormedSpace/HahnBanach/Extension.lean +++ b/Mathlib/Analysis/NormedSpace/HahnBanach/Extension.lean @@ -38,7 +38,9 @@ namespace Real variable {E : Type*} [SeminormedAddCommGroup E] [NormedSpace ℝ E] -/-- Hahn-Banach theorem for continuous linear functions over `ℝ`. -/ +/-- **Hahn-Banach theorem** for continuous linear functions over `ℝ`. +See also `exists_extension_norm_eq` in the root namespace for a more general version +that works both for `ℝ` and `ℂ`. -/ theorem exists_extension_norm_eq (p : Subspace ℝ E) (f : p →L[ℝ] ℝ) : ∃ g : E →L[ℝ] ℝ, (∀ x : p, g x = f x) ∧ ‖g‖ = ‖f‖ := by rcases exists_extension_of_le_sublinear ⟨p, f⟩ (fun x => ‖f‖ * ‖x‖) @@ -63,20 +65,22 @@ section IsROrC open IsROrC -variable {𝕜 : Type*} [IsROrC 𝕜] {F : Type*} [SeminormedAddCommGroup F] [NormedSpace 𝕜 F] +variable {𝕜 : Type*} [IsROrC 𝕜] {E F : Type*} + [SeminormedAddCommGroup E] [NormedSpace 𝕜 E] + [NormedAddCommGroup F] [NormedSpace 𝕜 F] -/-- Hahn-Banach theorem for continuous linear functions over `𝕜` satisfying `IsROrC 𝕜`. -/ -theorem exists_extension_norm_eq (p : Subspace 𝕜 F) (f : p →L[𝕜] 𝕜) : - ∃ g : F →L[𝕜] 𝕜, (∀ x : p, g x = f x) ∧ ‖g‖ = ‖f‖ := by - letI : Module ℝ F := RestrictScalars.module ℝ 𝕜 F - letI : IsScalarTower ℝ 𝕜 F := RestrictScalars.isScalarTower _ _ _ - letI : NormedSpace ℝ F := NormedSpace.restrictScalars _ 𝕜 _ +/-- **Hahn-Banach theorem** for continuous linear functions over `𝕜` satisfying `IsROrC 𝕜`. -/ +theorem exists_extension_norm_eq (p : Subspace 𝕜 E) (f : p →L[𝕜] 𝕜) : + ∃ g : E →L[𝕜] 𝕜, (∀ x : p, g x = f x) ∧ ‖g‖ = ‖f‖ := by + letI : Module ℝ E := RestrictScalars.module ℝ 𝕜 E + letI : IsScalarTower ℝ 𝕜 E := RestrictScalars.isScalarTower _ _ _ + letI : NormedSpace ℝ E := NormedSpace.restrictScalars _ 𝕜 _ -- Let `fr: p →L[ℝ] ℝ` be the real part of `f`. let fr := reClm.comp (f.restrictScalars ℝ) -- Use the real version to get a norm-preserving extension of `fr`, which - -- we'll call `g : F →L[ℝ] ℝ`. + -- we'll call `g : E →L[ℝ] ℝ`. rcases Real.exists_extension_norm_eq (p.restrictScalars ℝ) fr with ⟨g, ⟨hextends, hnormeq⟩⟩ - -- Now `g` can be extended to the `F →L[𝕜] 𝕜` we need. + -- Now `g` can be extended to the `E →L[𝕜] 𝕜` we need. refine' ⟨g.extendTo𝕜, _⟩ -- It is an extension of `f`. have h : ∀ x : p, g.extendTo𝕜 x = f x := by @@ -104,6 +108,34 @@ theorem exists_extension_norm_eq (p : Subspace 𝕜 F) (f : p →L[𝕜] 𝕜) : · exact f.op_norm_le_bound g.extendTo𝕜.op_norm_nonneg fun x => h x ▸ g.extendTo𝕜.le_op_norm x #align exists_extension_norm_eq exists_extension_norm_eq +open FiniteDimensional + +/-- Corollary of the **Hahn-Banach theorem**: if `f : p → F` is a continuous linear map +from a submodule of a normed space `E` over `𝕜`, `𝕜 = ℝ` or `𝕜 = ℂ`, +with a finite dimensional range, +then `f` admits an extension to a continuous linear map `E → F`. + +Note that contrary to the case `F = 𝕜`, see `exists_extension_norm_eq`, +we provide no estimates on the norm of the extension. +-/ +lemma ContinuousLinearMap.exist_extension_of_finiteDimensional_range {p : Submodule 𝕜 E} + (f : p →L[𝕜] F) [FiniteDimensional 𝕜 (LinearMap.range f)] : + ∃ g : E →L[𝕜] F, f = g.comp p.subtypeL := by + set b := finBasis 𝕜 (LinearMap.range f) + set e := b.equivFunL + set fi := fun i ↦ (LinearMap.toContinuousLinearMap (b.coord i)).comp + (f.codRestrict _ <| LinearMap.mem_range_self _) + choose gi hgf _ using fun i ↦ exists_extension_norm_eq p (fi i) + use (LinearMap.range f).subtypeL.comp <| e.symm.toContinuousLinearMap.comp (.pi gi) + ext x + simp [hgf] + +/-- A finite dimensional submodule over `ℝ` or `ℂ` is `Submodule.ClosedComplemented`. -/ +lemma Submodule.ClosedComplemented.of_finiteDimensional (p : Submodule 𝕜 F) + [FiniteDimensional 𝕜 p] : p.ClosedComplemented := + let ⟨g, hg⟩ := (ContinuousLinearMap.id 𝕜 p).exist_extension_of_finiteDimensional_range + ⟨g, FunLike.congr_fun hg.symm⟩ + end IsROrC section DualVector diff --git a/Mathlib/Analysis/NormedSpace/HahnBanach/Separation.lean b/Mathlib/Analysis/NormedSpace/HahnBanach/Separation.lean index 573c45b188845..d5951c64f9e09 100644 --- a/Mathlib/Analysis/NormedSpace/HahnBanach/Separation.lean +++ b/Mathlib/Analysis/NormedSpace/HahnBanach/Separation.lean @@ -55,7 +55,7 @@ theorem separate_convex_open_set [TopologicalSpace E] [AddCommGroup E] [Topologi rw [← f.domain.coe_mk x₀ (Submodule.mem_span_singleton_self _), hφ₁, LinearPMap.mkSpanSingleton'_apply_self] have hφ₄ : ∀ x ∈ s, φ x < 1 := fun x hx => - (hφ₂ x).trans_lt (gauge_lt_one_of_mem_of_open hs₂ hx) + (hφ₂ x).trans_lt (gauge_lt_one_of_mem_of_isOpen hs₂ hx) · refine' ⟨⟨φ, _⟩, hφ₃, hφ₄⟩ refine' φ.continuous_of_nonzero_on_open _ (hs₂.vadd (-x₀)) (Nonempty.vadd_set ⟨0, hs₀⟩) diff --git a/Mathlib/Analysis/NormedSpace/LpEquiv.lean b/Mathlib/Analysis/NormedSpace/LpEquiv.lean index 7a8eec5ee09b9..f6ea5c4a383a2 100644 --- a/Mathlib/Analysis/NormedSpace/LpEquiv.lean +++ b/Mathlib/Analysis/NormedSpace/LpEquiv.lean @@ -80,16 +80,16 @@ theorem equiv_lpPiLp_norm (f : lp E p) : ‖Equiv.lpPiLp f‖ = ‖f‖ := by #align equiv_lp_pi_Lp_norm equiv_lpPiLp_norm /-- The canonical `AddEquiv` between `lp E p` and `PiLp p E` when `E : α → Type u` with -`[Fintype α]` and `[Fact (1 ≤ p)]`. -/ -def AddEquiv.lpPiLp [Fact (1 ≤ p)] : lp E p ≃+ PiLp p E := +`[Fintype α]`. -/ +def AddEquiv.lpPiLp : lp E p ≃+ PiLp p E := { Equiv.lpPiLp with map_add' := fun _f _g => rfl } #align add_equiv.lp_pi_Lp AddEquiv.lpPiLp -theorem coe_addEquiv_lpPiLp [Fact (1 ≤ p)] (f : lp E p) : AddEquiv.lpPiLp f = ⇑f := +theorem coe_addEquiv_lpPiLp (f : lp E p) : AddEquiv.lpPiLp f = ⇑f := rfl #align coe_add_equiv_lp_pi_Lp coe_addEquiv_lpPiLp -theorem coe_addEquiv_lpPiLp_symm [Fact (1 ≤ p)] (f : PiLp p E) : +theorem coe_addEquiv_lpPiLp_symm (f : PiLp p E) : (AddEquiv.lpPiLp.symm f : ∀ i, E i) = f := rfl #align coe_add_equiv_lp_pi_Lp_symm coe_addEquiv_lpPiLp_symm diff --git a/Mathlib/Analysis/NormedSpace/MStructure.lean b/Mathlib/Analysis/NormedSpace/MStructure.lean index 63b83df1604ee..a1edf05ce370c 100644 --- a/Mathlib/Analysis/NormedSpace/MStructure.lean +++ b/Mathlib/Analysis/NormedSpace/MStructure.lean @@ -270,9 +270,9 @@ theorem distrib_lattice_lemma [FaithfulSMul M X] {P Q R : { P : M // IsLprojecti rw [add_mul, mul_add, mul_add, (mul_assoc _ (R : M) (↑Q * ↑R * ↑Pᶜ)), ← mul_assoc (R : M) (↑Q * ↑R) _, ← coe_inf Q, (Pᶜ.prop.commute R.prop).eq, ((Q ⊓ R).prop.commute Pᶜ.prop).eq, (R.prop.commute (Q ⊓ R).prop).eq, coe_inf Q, - mul_assoc (Q : M), ←mul_assoc, mul_assoc (R : M), (Pᶜ.prop.commute P.prop).eq, mul_compl_self, + mul_assoc (Q : M), ← mul_assoc, mul_assoc (R : M), (Pᶜ.prop.commute P.prop).eq, mul_compl_self, zero_mul, mul_zero, zero_add, add_zero, ← mul_assoc, P.prop.proj.eq, - R.prop.proj.eq, ←coe_inf Q, mul_assoc, ((Q ⊓ R).prop.commute Pᶜ.prop).eq, ← mul_assoc, + R.prop.proj.eq, ← coe_inf Q, mul_assoc, ((Q ⊓ R).prop.commute Pᶜ.prop).eq, ← mul_assoc, Pᶜ.prop.proj.eq] #align is_Lprojection.distrib_lattice_lemma IsLprojection.distrib_lattice_lemma diff --git a/Mathlib/Analysis/NormedSpace/Multilinear.lean b/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean similarity index 97% rename from Mathlib/Analysis/NormedSpace/Multilinear.lean rename to Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean index 783e622999185..55d3821e26a1d 100644 --- a/Mathlib/Analysis/NormedSpace/Multilinear.lean +++ b/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ import Mathlib.Analysis.NormedSpace.OperatorNorm -import Mathlib.Topology.Algebra.Module.Multilinear +import Mathlib.Topology.Algebra.Module.Multilinear.Basic #align_import analysis.normed_space.multilinear from "leanprover-community/mathlib"@"f40476639bac089693a489c9e354ebd75dc0f886" @@ -102,8 +102,9 @@ and elements `c i : 𝕜`, `1 < ‖c i‖`, then it satisfies this inequality fo theorem bound_of_shell {ε : ι → ℝ} {C : ℝ} (hε : ∀ i, 0 < ε i) {c : ι → 𝕜} (hc : ∀ i, 1 < ‖c i‖) (hf : ∀ m : ∀ i, E i, (∀ i, ε i / ‖c i‖ ≤ ‖m i‖) → (∀ i, ‖m i‖ < ε i) → ‖f m‖ ≤ C * ∏ i, ‖m i‖) (m : ∀ i, E i) : ‖f m‖ ≤ C * ∏ i, ‖m i‖ := by - rcases em (∃ i, m i = 0) with (⟨i, hi⟩ | hm) <;> [skip; push_neg at hm] + rcases em (∃ i, m i = 0) with (⟨i, hi⟩ | hm) · simp [f.map_coord_zero i hi, prod_eq_zero (mem_univ i), hi] + push_neg at hm choose δ hδ0 hδm_lt hle_δm _ using fun i => rescale_to_shell (hc i) (hε i) (hm i) have hδ0 : 0 < ∏ i, ‖δ i‖ := prod_pos fun i _ => norm_pos_iff.2 (hδ0 i) simpa [map_smul_univ, norm_smul, prod_mul_distrib, mul_left_comm C, mul_le_mul_left hδ0] using @@ -466,7 +467,8 @@ theorem isLeast_op_nnnorm : IsLeast {C : ℝ≥0 | ∀ m, ‖f m‖₊ ≤ C * theorem op_nnnorm_prod (f : ContinuousMultilinearMap 𝕜 E G) (g : ContinuousMultilinearMap 𝕜 E G') : ‖f.prod g‖₊ = max ‖f‖₊ ‖g‖₊ := - eq_of_forall_ge_iff fun _ ↦ by simp [op_nnnorm_le_iff, Prod.nnnorm_def', forall_and] + eq_of_forall_ge_iff fun _ ↦ by + simp only [op_nnnorm_le_iff, prod_apply, Prod.nnnorm_def', max_le_iff, forall_and] theorem op_norm_prod (f : ContinuousMultilinearMap 𝕜 E G) (g : ContinuousMultilinearMap 𝕜 E G') : ‖f.prod g‖ = max ‖f‖ ‖g‖ := @@ -575,12 +577,7 @@ def piₗᵢ {ι' : Type v'} [Fintype ι'] {E' : ι' → Type wE'} [∀ i', Norm @LinearIsometryEquiv 𝕜 𝕜 _ _ (RingHom.id 𝕜) _ _ _ (∀ i', ContinuousMultilinearMap 𝕜 E (E' i')) (ContinuousMultilinearMap 𝕜 E (∀ i, E' i)) _ _ (@Pi.module ι' _ 𝕜 _ _ fun _ => inferInstance) _ where - toLinearEquiv := - -- note: `piLinearEquiv` does not unify correctly here, - -- presumably due to issues with dependent typeclass arguments. - { piEquiv with - map_add' := fun _ _ => rfl - map_smul' := fun _ _ => rfl } + toLinearEquiv := piLinearEquiv norm_map' := op_norm_pi #align continuous_multilinear_map.piₗᵢ ContinuousMultilinearMap.piₗᵢ @@ -644,8 +641,8 @@ theorem norm_image_sub_le (m₁ m₂ : ∀ i, E i) : #align continuous_multilinear_map.norm_image_sub_le ContinuousMultilinearMap.norm_image_sub_le /-- Applying a multilinear map to a vector is continuous in both coordinates. -/ -theorem continuous_eval : Continuous fun p : - ContinuousMultilinearMap 𝕜 E G × ∀ i, E i => p.1 p.2 := by +theorem continuous_eval : Continuous + fun p : ContinuousMultilinearMap 𝕜 E G × ∀ i, E i => p.1 p.2 := by apply continuous_iff_continuousAt.2 fun p => ?_ apply continuousAt_of_locally_lipschitz zero_lt_one @@ -767,7 +764,7 @@ instance completeSpace [CompleteSpace G] : CompleteSpace (ContinuousMultilinearM have B : Tendsto (fun m => ‖(f n - f m) v‖) atTop (𝓝 ‖(f n - Fcont) v‖) := Tendsto.norm (tendsto_const_nhds.sub (hF v)) exact le_of_tendsto B A - erw [tendsto_iff_norm_sub_tendsto_zero] + rw [tendsto_iff_norm_sub_tendsto_zero] exact squeeze_zero (fun n => norm_nonneg _) this b_lim end ContinuousMultilinearMap @@ -814,9 +811,7 @@ variable {A : Type*} [NormedCommRing A] [NormedAlgebra 𝕜 A] @[simp] theorem norm_mkPiAlgebra_le [Nonempty ι] : ‖ContinuousMultilinearMap.mkPiAlgebra 𝕜 ι A‖ ≤ 1 := by - have := fun f => @op_norm_le_bound 𝕜 ι (fun _ => A) A _ _ _ _ _ _ f _ zero_le_one - refine' this _ _ - intro m + refine op_norm_le_bound _ zero_le_one fun m => ?_ simp only [ContinuousMultilinearMap.mkPiAlgebra_apply, one_mul] exact norm_prod_le' _ univ_nonempty _ #align continuous_multilinear_map.norm_mk_pi_algebra_le ContinuousMultilinearMap.norm_mkPiAlgebra_le @@ -824,11 +819,9 @@ theorem norm_mkPiAlgebra_le [Nonempty ι] : ‖ContinuousMultilinearMap.mkPiAlge theorem norm_mkPiAlgebra_of_empty [IsEmpty ι] : ‖ContinuousMultilinearMap.mkPiAlgebra 𝕜 ι A‖ = ‖(1 : A)‖ := by apply le_antisymm - · have := fun f => @op_norm_le_bound 𝕜 ι (fun _ => A) A _ _ _ _ _ _ f _ (norm_nonneg (1 : A)) - refine' this _ _ - simp - -- Porting note: have to annotate types to get mvars to unify - · convert ratio_le_op_norm (ContinuousMultilinearMap.mkPiAlgebra 𝕜 ι A) fun _ => (1 : A) + · apply op_norm_le_bound <;> simp + · -- Porting note: have to annotate types to get mvars to unify + convert ratio_le_op_norm (ContinuousMultilinearMap.mkPiAlgebra 𝕜 ι A) fun _ => (1 : A) simp [eq_empty_of_isEmpty (univ : Finset ι)] #align continuous_multilinear_map.norm_mk_pi_algebra_of_empty ContinuousMultilinearMap.norm_mkPiAlgebra_of_empty @@ -848,16 +841,13 @@ section variable {A : Type*} [NormedRing A] [NormedAlgebra 𝕜 A] theorem norm_mkPiAlgebraFin_succ_le : ‖ContinuousMultilinearMap.mkPiAlgebraFin 𝕜 n.succ A‖ ≤ 1 := by - have := fun f => @op_norm_le_bound 𝕜 (Fin n.succ) (fun _ => A) A _ _ _ _ _ _ f _ zero_le_one - refine' this _ _ - intro m + refine op_norm_le_bound _ zero_le_one fun m => ?_ simp only [ContinuousMultilinearMap.mkPiAlgebraFin_apply, one_mul, List.ofFn_eq_map, Fin.prod_univ_def, Multiset.coe_map, Multiset.coe_prod] refine' (List.norm_prod_le' _).trans_eq _ · rw [Ne.def, List.map_eq_nil, List.finRange_eq_nil] exact Nat.succ_ne_zero _ - rw [List.map_map] - rfl + rw [List.map_map, Function.comp_def] #align continuous_multilinear_map.norm_mk_pi_algebra_fin_succ_le ContinuousMultilinearMap.norm_mkPiAlgebraFin_succ_le theorem norm_mkPiAlgebraFin_le_of_pos (hn : 0 < n) : @@ -868,9 +858,7 @@ theorem norm_mkPiAlgebraFin_le_of_pos (hn : 0 < n) : theorem norm_mkPiAlgebraFin_zero : ‖ContinuousMultilinearMap.mkPiAlgebraFin 𝕜 0 A‖ = ‖(1 : A)‖ := by refine' le_antisymm _ _ - · have := fun f => - @op_norm_le_bound 𝕜 (Fin 0) (fun _ => A) A _ _ _ _ _ _ f _ (norm_nonneg (1 : A)) - refine' this _ _ + · refine op_norm_le_bound _ (norm_nonneg (1 : A)) ?_ simp · convert ratio_le_op_norm (ContinuousMultilinearMap.mkPiAlgebraFin 𝕜 0 A) fun _ => (1 : A) simp @@ -883,8 +871,8 @@ theorem norm_mkPiAlgebraFin [NormOneClass A] : · rw [norm_mkPiAlgebraFin_zero] simp · refine' le_antisymm norm_mkPiAlgebraFin_succ_le _ - convert ratio_le_op_norm (ContinuousMultilinearMap.mkPiAlgebraFin 𝕜 (Nat.succ _) A) - fun _ => 1 + refine le_of_eq_of_le ?_ <| + ratio_le_op_norm (ContinuousMultilinearMap.mkPiAlgebraFin 𝕜 (Nat.succ _) A) fun _ => 1 simp #align continuous_multilinear_map.norm_mk_pi_algebra_fin ContinuousMultilinearMap.norm_mkPiAlgebraFin @@ -896,8 +884,7 @@ variable (𝕜 ι) `m i` (multiplied by a fixed reference element `z` in the target module) -/ protected def mkPiField (z : G) : ContinuousMultilinearMap 𝕜 (fun _ : ι => 𝕜) G := MultilinearMap.mkContinuous (MultilinearMap.mkPiRing 𝕜 ι z) ‖z‖ fun m => by - simp only [MultilinearMap.mkPiRing_apply, norm_smul, norm_prod, mul_comm] - rfl + simp only [MultilinearMap.mkPiRing_apply, norm_smul, norm_prod, mul_comm, le_rfl] #align continuous_multilinear_map.mk_pi_field ContinuousMultilinearMap.mkPiField variable {𝕜 ι} @@ -975,15 +962,10 @@ def compContinuousMultilinearMapL : (G →L[𝕜] G') →L[𝕜] ContinuousMultilinearMap 𝕜 E G →L[𝕜] ContinuousMultilinearMap 𝕜 E G' := LinearMap.mkContinuous₂ (LinearMap.mk₂ 𝕜 compContinuousMultilinearMap (fun f₁ f₂ g => rfl) (fun c f g => rfl) - (fun f g₁ g₂ => by - ext1 - apply f.map_add) - fun c f g => by - ext1 - simp) - 1 fun f g => by - rw [one_mul] - exact f.norm_compContinuousMultilinearMap_le g + (fun f g₁ g₂ => by ext1; apply f.map_add) + (fun c f g => by ext1; simp)) + 1 + fun f g => by rw [one_mul]; exact f.norm_compContinuousMultilinearMap_le g #align continuous_linear_map.comp_continuous_multilinear_mapL ContinuousLinearMap.compContinuousMultilinearMapL variable {𝕜 G G'} @@ -993,8 +975,7 @@ continuous linear equiv. -/ nonrec def _root_.ContinuousLinearEquiv.compContinuousMultilinearMapL (g : G ≃L[𝕜] G') : ContinuousMultilinearMap 𝕜 E G ≃L[𝕜] ContinuousMultilinearMap 𝕜 E G' := - { compContinuousMultilinearMapL 𝕜 E G G' - g.toContinuousLinearMap with + { compContinuousMultilinearMapL 𝕜 E G G' g.toContinuousLinearMap with invFun := compContinuousMultilinearMapL 𝕜 E G' G g.symm.toContinuousLinearMap left_inv := by intro f @@ -1170,7 +1151,7 @@ theorem norm_compContinuous_linearIsometry_le (g : ContinuousMultilinearMap 𝕜 refine op_norm_le_bound _ (norm_nonneg _) fun m => ?_ apply (g.le_op_norm _).trans _ simp only [ContinuousLinearMap.coe_coe, LinearIsometry.coe_toContinuousLinearMap, - LinearIsometry.norm_map]; rfl + LinearIsometry.norm_map, le_rfl] #align continuous_multilinear_map.norm_comp_continuous_linear_isometry_le ContinuousMultilinearMap.norm_compContinuous_linearIsometry_le theorem norm_compContinuous_linearIsometryEquiv (g : ContinuousMultilinearMap 𝕜 E₁ G) @@ -1198,8 +1179,9 @@ def compContinuousLinearMapL (f : ∀ i, E i →L[𝕜] E₁ i) : LinearMap.mkContinuous { toFun := fun g => g.compContinuousLinearMap f map_add' := fun _ _ => rfl - map_smul' := fun _ _ => rfl } (∏ i, ‖f i‖) fun _ => - (norm_compContinuousLinearMap_le _ _).trans_eq (mul_comm _ _) + map_smul' := fun _ _ => rfl } + (∏ i, ‖f i‖) + fun _ => (norm_compContinuousLinearMap_le _ _).trans_eq (mul_comm _ _) #align continuous_multilinear_map.comp_continuous_linear_mapL ContinuousMultilinearMap.compContinuousLinearMapL @[simp] @@ -1211,7 +1193,7 @@ theorem compContinuousLinearMapL_apply (g : ContinuousMultilinearMap 𝕜 E₁ G variable (G) theorem norm_compContinuousLinearMapL_le (f : ∀ i, E i →L[𝕜] E₁ i) : - ‖@compContinuousLinearMapL 𝕜 ι E E₁ G _ _ _ _ _ _ _ _ f‖ ≤ ∏ i, ‖f i‖ := + ‖compContinuousLinearMapL (G := G) f‖ ≤ ∏ i, ‖f i‖ := LinearMap.mkContinuous_norm_le _ (prod_nonneg fun _ _ => norm_nonneg _) _ #align continuous_multilinear_map.norm_comp_continuous_linear_mapL_le ContinuousMultilinearMap.norm_compContinuousLinearMapL_le @@ -1296,7 +1278,7 @@ theorem compContinuousLinearMapEquivL_apply (g : ContinuousMultilinearMap 𝕜 E end ContinuousMultilinearMap -section Smul +section SMul variable {R : Type*} [Semiring R] [Module R G] [SMulCommClass 𝕜 R G] [ContinuousConstSMul R G] @@ -1304,4 +1286,4 @@ instance continuousConstSMul : ContinuousConstSMul R (ContinuousMultilinearMap ⟨fun c => (ContinuousLinearMap.compContinuousMultilinearMapL 𝕜 _ G G (c • ContinuousLinearMap.id 𝕜 G)).2⟩ -end Smul +end SMul diff --git a/Mathlib/Analysis/NormedSpace/MultilinearCurrying.lean b/Mathlib/Analysis/NormedSpace/Multilinear/Curry.lean similarity index 97% rename from Mathlib/Analysis/NormedSpace/MultilinearCurrying.lean rename to Mathlib/Analysis/NormedSpace/Multilinear/Curry.lean index c17ed7aa9f1dc..a2b6e5e3e47d3 100644 --- a/Mathlib/Analysis/NormedSpace/MultilinearCurrying.lean +++ b/Mathlib/Analysis/NormedSpace/Multilinear/Curry.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ -import Mathlib.Analysis.NormedSpace.Multilinear +import Mathlib.Analysis.NormedSpace.Multilinear.Basic #align_import analysis.normed_space.multilinear from "leanprover-community/mathlib"@"f40476639bac089693a489c9e354ebd75dc0f886" @@ -157,9 +157,8 @@ theorem ContinuousLinearMap.curry_uncurryLeft (f : Ei 0 →L[𝕜] ContinuousMultilinearMap 𝕜 (fun i : Fin n => Ei i.succ) G) : f.uncurryLeft.curryLeft = f := by ext m x - simp only [tail_cons, ContinuousLinearMap.uncurryLeft_apply, - ContinuousMultilinearMap.curryLeft_apply] - rw [cons_zero] + rw [ContinuousMultilinearMap.curryLeft_apply, ContinuousLinearMap.uncurryLeft_apply, tail_cons, + cons_zero] #align continuous_linear_map.curry_uncurry_left ContinuousLinearMap.curry_uncurryLeft @[simp] @@ -191,11 +190,13 @@ def continuousMultilinearCurryLeftEquiv : rfl invFun := ContinuousMultilinearMap.curryLeft left_inv := ContinuousLinearMap.curry_uncurryLeft - right_inv := ContinuousMultilinearMap.uncurry_curryLeft } (fun f => by - simp only [LinearEquiv.coe_mk] - exact MultilinearMap.mkContinuous_norm_le _ (norm_nonneg f) _) fun f => by - simp only [LinearEquiv.coe_symm_mk] - exact LinearMap.mkContinuous_norm_le _ (norm_nonneg f) _ + right_inv := ContinuousMultilinearMap.uncurry_curryLeft } + (fun f => by + simp only [LinearEquiv.coe_mk] + exact MultilinearMap.mkContinuous_norm_le _ (norm_nonneg f) _) + (fun f => by + simp only [LinearEquiv.coe_symm_mk] + exact LinearMap.mkContinuous_norm_le _ (norm_nonneg f) _) #align continuous_multilinear_curry_left_equiv continuousMultilinearCurryLeftEquiv variable {𝕜 Ei G} @@ -261,11 +262,11 @@ def ContinuousMultilinearMap.curryRight (f : ContinuousMultilinearMap 𝕜 Ei G) (f.toMultilinearMap.curryRight m).mkContinuous (‖f‖ * ∏ i, ‖m i‖) fun x => f.norm_map_snoc_le m x map_add' := fun m i x y => by + ext simp - rfl map_smul' := fun m i c x => by - simp - rfl } + ext + simp } f'.mkContinuous ‖f‖ fun m => by simp only [MultilinearMap.coe_mk] exact LinearMap.mkContinuous_norm_le _ @@ -283,16 +284,15 @@ theorem ContinuousMultilinearMap.curry_uncurryRight (f : ContinuousMultilinearMap 𝕜 (fun i : Fin n => Ei <| castSucc i) (Ei (last n) →L[𝕜] G)) : f.uncurryRight.curryRight = f := by ext m x - simp only [snoc_last, ContinuousMultilinearMap.curryRight_apply, - ContinuousMultilinearMap.uncurryRight_apply] - rw [init_snoc] + rw [ContinuousMultilinearMap.curryRight_apply, ContinuousMultilinearMap.uncurryRight_apply, + snoc_last, init_snoc] #align continuous_multilinear_map.curry_uncurry_right ContinuousMultilinearMap.curry_uncurryRight @[simp] theorem ContinuousMultilinearMap.uncurry_curryRight (f : ContinuousMultilinearMap 𝕜 Ei G) : f.curryRight.uncurryRight = f := by ext m - simp + rw [uncurryRight_apply, curryRight_apply, snoc_init_self] #align continuous_multilinear_map.uncurry_curry_right ContinuousMultilinearMap.uncurry_curryRight variable (𝕜 Ei G) @@ -430,7 +430,7 @@ theorem ContinuousMultilinearMap.uncurry0_apply (f : G[×0]→L[𝕜] G') : f.un theorem ContinuousMultilinearMap.apply_zero_curry0 (f : G[×0]→L[𝕜] G') {x : Fin 0 → G} : ContinuousMultilinearMap.curry0 𝕜 G (f x) = f := by ext m - simp [(Subsingleton.elim _ _ : x = m)] + simp [Subsingleton.elim x m] #align continuous_multilinear_map.apply_zero_curry0 ContinuousMultilinearMap.apply_zero_curry0 theorem ContinuousMultilinearMap.uncurry0_curry0 (f : G[×0]→L[𝕜] G') : @@ -670,7 +670,6 @@ theorem curryFinFinset_apply_const (hk : s.card = k) (hl : sᶜ.card = l) (f : G (x y : G) : (curryFinFinset 𝕜 G G' hk hl f (fun _ => x) fun _ => y) = f (s.piecewise (fun _ => x) fun _ => y) := by refine' (curryFinFinset_symm_apply_piecewise_const hk hl _ _ _).symm.trans _ - -- `rw` fails rw [LinearIsometryEquiv.symm_apply_apply] #align continuous_multilinear_map.curry_fin_finset_apply_const ContinuousMultilinearMap.curryFinFinset_apply_const diff --git a/Mathlib/Analysis/NormedSpace/OperatorNorm.lean b/Mathlib/Analysis/NormedSpace/OperatorNorm.lean index eaa369ff76005..b67899fbda8c3 100644 --- a/Mathlib/Analysis/NormedSpace/OperatorNorm.lean +++ b/Mathlib/Analysis/NormedSpace/OperatorNorm.lean @@ -29,7 +29,9 @@ is isometric, as expressed by the typeclass `[RingHomIsometric σ]`. suppress_compilation -open Classical NNReal Topology Bornology +open Bornology +open Filter hiding map_smul +open scoped Classical NNReal Topology Uniformity -- the `ₗ` subscript variables are for special cases about linear (as opposed to semilinear) maps variable {𝕜 𝕜₂ 𝕜₃ E Eₗ F Fₗ G Gₗ 𝓕 : Type*} @@ -291,108 +293,53 @@ theorem op_norm_smul_le {𝕜' : Type*} [NormedField 𝕜'] [NormedSpace 𝕜' F exact mul_le_mul_of_nonneg_left (le_op_norm _ _) (norm_nonneg _) #align continuous_linear_map.op_norm_smul_le ContinuousLinearMap.op_norm_smul_le -/-- Continuous linear maps themselves form a seminormed space with respect to -the operator norm. This is only a temporary definition because we want to replace the topology -with `ContinuousLinearMap.topologicalSpace` to avoid diamond issues. -See Note [forgetful inheritance] -/ -protected def tmpSeminormedAddCommGroup : SeminormedAddCommGroup (E →SL[σ₁₂] F) := - AddGroupSeminorm.toSeminormedAddCommGroup - { toFun := norm - map_zero' := op_norm_zero - add_le' := op_norm_add_le - neg' := op_norm_neg } -#align continuous_linear_map.tmp_seminormed_add_comm_group ContinuousLinearMap.tmpSeminormedAddCommGroup - -/-- The `PseudoMetricSpace` structure on `E →SL[σ₁₂] F` coming from -`ContinuousLinearMap.tmpSeminormedAddCommGroup`. -See Note [forgetful inheritance] -/ -protected def tmpPseudoMetricSpace : PseudoMetricSpace (E →SL[σ₁₂] F) := - ContinuousLinearMap.tmpSeminormedAddCommGroup.toPseudoMetricSpace -#align continuous_linear_map.tmp_pseudo_metric_space ContinuousLinearMap.tmpPseudoMetricSpace - -/-- The `UniformSpace` structure on `E →SL[σ₁₂] F` coming from -`ContinuousLinearMap.tmpSeminormedAddCommGroup`. -See Note [forgetful inheritance] -/ -protected def tmpUniformSpace : UniformSpace (E →SL[σ₁₂] F) := - ContinuousLinearMap.tmpPseudoMetricSpace.toUniformSpace -#align continuous_linear_map.tmp_uniform_space ContinuousLinearMap.tmpUniformSpace - -/-- The `TopologicalSpace` structure on `E →SL[σ₁₂] F` coming from -`ContinuousLinearMap.tmpSeminormedAddCommGroup`. -See Note [forgetful inheritance] -/ -protected def tmpTopologicalSpace : TopologicalSpace (E →SL[σ₁₂] F) := - ContinuousLinearMap.tmpUniformSpace.toTopologicalSpace -#align continuous_linear_map.tmp_topological_space ContinuousLinearMap.tmpTopologicalSpace - -section Tmp - -attribute [-instance] ContinuousLinearMap.topologicalSpace - -attribute [-instance] ContinuousLinearMap.uniformSpace - -attribute [local instance] ContinuousLinearMap.tmpSeminormedAddCommGroup - -protected theorem tmpTopologicalAddGroup : TopologicalAddGroup (E →SL[σ₁₂] F) := - inferInstance -#align continuous_linear_map.tmp_topological_add_group ContinuousLinearMap.tmpTopologicalAddGroup - -protected theorem tmp_closedBall_div_subset {a b : ℝ} (ha : 0 < a) (hb : 0 < b) : - closedBall (0 : E →SL[σ₁₂] F) (a / b) ⊆ - { f | ∀ x ∈ closedBall (0 : E) b, f x ∈ closedBall (0 : F) a } := by - intro f hf x hx - rw [mem_closedBall_zero_iff] at hf hx ⊢ - calc - ‖f x‖ ≤ ‖f‖ * ‖x‖ := le_op_norm _ _ - _ ≤ a / b * b := by gcongr - _ = a := div_mul_cancel a hb.ne.symm -#align continuous_linear_map.tmp_closed_ball_div_subset ContinuousLinearMap.tmp_closedBall_div_subset - -end Tmp - -protected theorem tmp_topology_eq : - (ContinuousLinearMap.tmpTopologicalSpace : TopologicalSpace (E →SL[σ₁₂] F)) = - ContinuousLinearMap.topologicalSpace := by - refine' - ContinuousLinearMap.tmpTopologicalAddGroup.ext inferInstance - ((@Metric.nhds_basis_closedBall _ ContinuousLinearMap.tmpPseudoMetricSpace 0).ext - (ContinuousLinearMap.hasBasis_nhds_zero_of_basis Metric.nhds_basis_closedBall) _ _) - · rcases NormedField.exists_norm_lt_one 𝕜 with ⟨c, hc₀, hc₁⟩ - intro ε hε - refine' ⟨⟨closedBall 0 (1 / ‖c‖), ε⟩, ⟨⟨_, hε⟩, _⟩⟩ - · exact NormedSpace.isVonNBounded_closedBall _ _ _ - intro f (hf : ∀ x, _) - simp_rw [mem_closedBall_zero_iff] at hf - convert (@mem_closedBall_zero_iff _ (_) f ε).2 _ -- Porting note: needed `convert` - refine' op_norm_le_of_shell' (div_pos one_pos hc₀) hε.le hc₁ fun x hx₁ hxc => _ - rw [div_mul_cancel 1 hc₀.ne.symm] at hx₁ - exact (hf x hxc.le).trans (le_mul_of_one_le_right hε.le hx₁) - · rintro ⟨S, ε⟩ ⟨hS, hε⟩ - rw [NormedSpace.isVonNBounded_iff] at hS - rcases hS.subset_closedBall_lt 0 0 with ⟨δ, hδ, hSδ⟩ - exact ⟨ε / δ, div_pos hε hδ, - (ContinuousLinearMap.tmp_closedBall_div_subset hε hδ).trans fun f hf x hx => hf x <| hSδ hx⟩ -#align continuous_linear_map.tmp_topology_eq ContinuousLinearMap.tmp_topology_eq - -protected theorem tmpUniformSpace_eq : - (ContinuousLinearMap.tmpUniformSpace : UniformSpace (E →SL[σ₁₂] F)) = - ContinuousLinearMap.uniformSpace := by - rw [← @UniformAddGroup.toUniformSpace_eq _ ContinuousLinearMap.tmpUniformSpace, ← - @UniformAddGroup.toUniformSpace_eq _ ContinuousLinearMap.uniformSpace] - congr! 1 - exact ContinuousLinearMap.tmp_topology_eq -#align continuous_linear_map.tmp_uniform_space_eq ContinuousLinearMap.tmpUniformSpace_eq - -instance toPseudoMetricSpace : PseudoMetricSpace (E →SL[σ₁₂] F) := - ContinuousLinearMap.tmpPseudoMetricSpace.replaceUniformity - (congr_arg _ ContinuousLinearMap.tmpUniformSpace_eq.symm) +/-- Operator seminorm on the space of continuous (semi)linear maps, as `Seminorm`. + +We use this seminorm to define a `SeminormedGroup` structure on `E →SL[σ] F`, +but we have to override the projection `UniformSpace` +so that it is definitionally equal to the one coming from the topologies on `E` and `F`. -/ +protected def seminorm : Seminorm 𝕜₂ (E →SL[σ₁₂] F) := + .ofSMulLE norm op_norm_zero op_norm_add_le op_norm_smul_le + +private lemma uniformity_eq_seminorm : + 𝓤 (E →SL[σ₁₂] F) = ⨅ r > 0, 𝓟 {f | ‖f.1 - f.2‖ < r} := by + refine ContinuousLinearMap.seminorm.uniformity_eq_of_hasBasis + (ContinuousLinearMap.hasBasis_nhds_zero_of_basis Metric.nhds_basis_closedBall) + ?_ fun (s, r) ⟨hs, hr⟩ ↦ ?_ + · rcases NormedField.exists_lt_norm 𝕜 1 with ⟨c, hc⟩ + refine ⟨‖c‖, ContinuousLinearMap.hasBasis_nhds_zero.mem_iff.2 + ⟨(closedBall 0 1, closedBall 0 1), ?_⟩⟩ + suffices ∀ f : E →SL[σ₁₂] F, (∀ x, ‖x‖ ≤ 1 → ‖f x‖ ≤ 1) → ‖f‖ ≤ ‖c‖ by + simpa [NormedSpace.isVonNBounded_closedBall, closedBall_mem_nhds, subset_def] using this + intro f hf + refine op_norm_le_of_shell (f := f) one_pos (norm_nonneg c) hc fun x hcx hx ↦ ?_ + exact (hf x hx.le).trans ((div_le_iff' <| one_pos.trans hc).1 hcx) + · rcases (NormedSpace.isVonNBounded_iff' _ _ _).1 hs with ⟨ε, hε⟩ + rcases exists_pos_mul_lt hr ε with ⟨δ, hδ₀, hδ⟩ + refine ⟨δ, hδ₀, fun f hf x hx ↦ ?_⟩ + simp only [Seminorm.mem_ball_zero, mem_closedBall_zero_iff] at hf ⊢ + rw [mul_comm] at hδ + exact le_trans (le_of_op_norm_le_of_le _ hf.le (hε _ hx)) hδ.le + +instance toPseudoMetricSpace : PseudoMetricSpace (E →SL[σ₁₂] F) := .replaceUniformity + ContinuousLinearMap.seminorm.toSeminormedAddCommGroup.toPseudoMetricSpace uniformity_eq_seminorm #align continuous_linear_map.to_pseudo_metric_space ContinuousLinearMap.toPseudoMetricSpace /-- Continuous linear maps themselves form a seminormed space with respect to the operator norm. -/ instance toSeminormedAddCommGroup : SeminormedAddCommGroup (E →SL[σ₁₂] F) where - dist_eq := ContinuousLinearMap.tmpSeminormedAddCommGroup.dist_eq + dist_eq _ _ := rfl #align continuous_linear_map.to_seminormed_add_comm_group ContinuousLinearMap.toSeminormedAddCommGroup +#noalign continuous_linear_map.tmp_seminormed_add_comm_group +#noalign continuous_linear_map.tmp_pseudo_metric_space +#noalign continuous_linear_map.tmp_uniform_space +#noalign continuous_linear_map.tmp_topological_space +#noalign continuous_linear_map.tmp_topological_add_group +#noalign continuous_linear_map.tmp_closed_ball_div_subset +#noalign continuous_linear_map.tmp_topology_eq +#noalign continuous_linear_map.tmp_uniform_space_eq + theorem nnnorm_def (f : E →SL[σ₁₂] F) : ‖f‖₊ = sInf { c | ∀ x, ‖f x‖₊ ≤ c * ‖x‖₊ } := by ext rw [NNReal.coe_sInf, coe_nnnorm, norm_def, NNReal.coe_image] @@ -1436,8 +1383,7 @@ theorem bound_of_shell [RingHomIsometric σ₁₂] (f : E →ₛₗ[σ₁₂] F) (hc : 1 < ‖c‖) (hf : ∀ x, ε / ‖c‖ ≤ ‖x‖ → ‖x‖ < ε → ‖f x‖ ≤ C * ‖x‖) (x : E) : ‖f x‖ ≤ C * ‖x‖ := by by_cases hx : x = 0; · simp [hx] - exact - SemilinearMapClass.bound_of_shell_semi_normed f ε_pos hc hf (ne_of_lt (norm_pos_iff.2 hx)).symm + exact SemilinearMapClass.bound_of_shell_semi_normed f ε_pos hc hf (norm_ne_zero_iff.2 hx) #align linear_map.bound_of_shell LinearMap.bound_of_shell /-- `LinearMap.bound_of_ball_bound'` is a version of this lemma over a field satisfying `IsROrC` diff --git a/Mathlib/Analysis/NormedSpace/RieszLemma.lean b/Mathlib/Analysis/NormedSpace/RieszLemma.lean index c11d8079f9662..564a0d5dac40c 100644 --- a/Mathlib/Analysis/NormedSpace/RieszLemma.lean +++ b/Mathlib/Analysis/NormedSpace/RieszLemma.lean @@ -104,7 +104,7 @@ theorem riesz_lemma_of_norm_lt {c : 𝕜} (hc : 1 < ‖c‖) {R : ℝ} (hR : ‖ simp only [norm_smul] ring _ ≤ ‖d‖ * ‖x - y'‖ := by gcongr; exact hx y' (by simp [Submodule.smul_mem _ _ hy]) - _ = ‖d • x - y‖ := by rw [yy', ←smul_sub, norm_smul] + _ = ‖d • x - y‖ := by rw [yy', ← smul_sub, norm_smul] #align riesz_lemma_of_norm_lt riesz_lemma_of_norm_lt theorem Metric.closedBall_infDist_compl_subset_closure {x : F} {s : Set F} (hx : x ∈ s) : diff --git a/Mathlib/Analysis/NormedSpace/Spectrum.lean b/Mathlib/Analysis/NormedSpace/Spectrum.lean index 5b42fbed98434..1ced47bd71136 100644 --- a/Mathlib/Analysis/NormedSpace/Spectrum.lean +++ b/Mathlib/Analysis/NormedSpace/Spectrum.lean @@ -226,7 +226,7 @@ theorem hasDerivAt_resolvent {a : A} {k : 𝕜} (hk : k ∈ ρ a) : #noalign spectrum.norm_resolvent_le_forall theorem eventually_isUnit_resolvent (a : A) : ∀ᶠ z in cobounded 𝕜, IsUnit (resolvent a z) := by - rw [←comap_norm_atTop, atTop_basis.comap (‖·‖) |>.eventually_iff] + rw [← comap_norm_atTop, atTop_basis.comap (‖·‖) |>.eventually_iff] refine ⟨‖a‖ * ‖(1 : A)‖ + 1, by trivial, fun z hz ↦ ?_⟩ exact isUnit_resolvent.mp <| mem_resolventSet_of_norm_lt_mul <| (lt_add_one (‖a‖ * _)).trans_le hz @@ -380,7 +380,7 @@ variable [NormedRing A] [NormedAlgebra ℂ A] [CompleteSpace A] [Nontrivial A] ( protected theorem nonempty : (spectrum ℂ a).Nonempty := by /- Suppose `σ a = ∅`, then resolvent set is `ℂ`, any `(z • 1 - a)` is a unit, and `resolvent a` is differentiable on `ℂ`. -/ - by_contra' h + by_contra! h have H₀ : resolventSet ℂ a = Set.univ := by rwa [spectrum, Set.compl_empty_iff] at h have H₁ : Differentiable ℂ fun z : ℂ => resolvent a z := fun z => (hasDerivAt_resolvent (H₀.symm ▸ Set.mem_univ z : z ∈ resolventSet ℂ a)).differentiableAt diff --git a/Mathlib/Analysis/NormedSpace/Star/Unitization.lean b/Mathlib/Analysis/NormedSpace/Star/Unitization.lean index d429a9b875203..edf38649388da 100644 --- a/Mathlib/Analysis/NormedSpace/Star/Unitization.lean +++ b/Mathlib/Analysis/NormedSpace/Star/Unitization.lean @@ -102,7 +102,7 @@ theorem Unitization.norm_splitMul_snd_sq (x : Unitization 𝕜 E) : refine (norm_mul_le _ _).trans ?_ calc _ ≤ ‖star x.fst • (x.fst • b + x.snd * b) + star x.snd * (x.fst • b + x.snd * b)‖ := by - nth_rewrite 2 [←one_mul ‖_ + _‖] + nth_rewrite 2 [← one_mul ‖_ + _‖] gcongr exact (norm_star b).symm ▸ mem_closedBall_zero_iff.1 hb _ ≤ sSup (_ '' Metric.closedBall 0 1) := le_csSup ?_ ⟨b, hb, ?_⟩ @@ -133,7 +133,7 @@ instance Unitization.instCstarRing : CstarRing (Unitization 𝕜 E) where ‖(Unitization.splitMul 𝕜 E x).snd‖ ≤ ‖(Unitization.splitMul 𝕜 E (star x)).snd‖ := by simp only [add_zero, Unitization.splitMul_apply, Unitization.snd_star, Unitization.fst_star] intro x - /- split based on whether the term inside the nprm is zero or not. If so, it's trivial. + /- split based on whether the term inside the norm is zero or not. If so, it's trivial. If not, then apply `norm_splitMul_snd_sq` and cancel one copy of the norm -/ by_cases h : algebraMap 𝕜 (E →L[𝕜] E) x.fst + mul 𝕜 E x.snd = 0 · simp only [h, norm_zero, norm_le_zero_iff] diff --git a/Mathlib/Analysis/Seminorm.lean b/Mathlib/Analysis/Seminorm.lean index 09883f6bd7f20..bc2ebe489998b 100644 --- a/Mathlib/Analysis/Seminorm.lean +++ b/Mathlib/Analysis/Seminorm.lean @@ -40,7 +40,7 @@ set_option autoImplicit true open NormedField Set Filter -open BigOperators NNReal Pointwise Topology +open scoped BigOperators NNReal Pointwise Topology Uniformity variable {R R' 𝕜 𝕜₂ 𝕜₃ 𝕝 E E₂ E₃ F G ι : Type*} @@ -396,7 +396,7 @@ theorem finset_sup_apply (p : ι → Seminorm 𝕜 E) (s : Finset ι) (x : E) : · rw [Finset.sup_empty, Finset.sup_empty, coe_bot, _root_.bot_eq_zero, Pi.zero_apply] norm_cast · rw [Finset.sup_cons, Finset.sup_cons, coe_sup, sup_eq_max, Pi.sup_apply, sup_eq_max, - NNReal.coe_max, coe_mk, ih] + NNReal.coe_max, NNReal.coe_mk, ih] #align seminorm.finset_sup_apply Seminorm.finset_sup_apply theorem exists_apply_eq_finset_sup (p : ι → Seminorm 𝕜 E) {s : Finset ι} (hs : s.Nonempty) (x : E) : @@ -897,14 +897,14 @@ theorem ball_finset_sup_eq_iInter (p : ι → Seminorm 𝕜 E) (s : Finset ι) ( (hr : 0 < r) : ball (s.sup p) x r = ⋂ i ∈ s, ball (p i) x r := by lift r to NNReal using hr.le simp_rw [ball, iInter_setOf, finset_sup_apply, NNReal.coe_lt_coe, - Finset.sup_lt_iff (show ⊥ < r from hr), ← NNReal.coe_lt_coe, coe_mk] + Finset.sup_lt_iff (show ⊥ < r from hr), ← NNReal.coe_lt_coe, NNReal.coe_mk] #align seminorm.ball_finset_sup_eq_Inter Seminorm.ball_finset_sup_eq_iInter theorem closedBall_finset_sup_eq_iInter (p : ι → Seminorm 𝕜 E) (s : Finset ι) (x : E) {r : ℝ} (hr : 0 ≤ r) : closedBall (s.sup p) x r = ⋂ i ∈ s, closedBall (p i) x r := by lift r to NNReal using hr simp_rw [closedBall, iInter_setOf, finset_sup_apply, NNReal.coe_le_coe, Finset.sup_le_iff, ← - NNReal.coe_le_coe, coe_mk] + NNReal.coe_le_coe, NNReal.coe_mk] #align seminorm.closed_ball_finset_sup_eq_Inter Seminorm.closedBall_finset_sup_eq_iInter theorem ball_finset_sup (p : ι → Seminorm 𝕜 E) (s : Finset ι) (x : E) {r : ℝ} (hr : 0 < r) : @@ -1175,19 +1175,13 @@ theorem continuousAt_zero_of_forall' [TopologicalSpace E] {p : Seminorm 𝕝 E} theorem continuousAt_zero' [TopologicalSpace E] [ContinuousConstSMul 𝕜 E] {p : Seminorm 𝕜 E} {r : ℝ} (hp : p.closedBall 0 r ∈ (𝓝 0 : Filter E)) : ContinuousAt p 0 := by - let r' := max 1 r - have hr' : 0 < r' := lt_max_of_lt_left one_pos - have hp' : p.closedBall 0 r' ∈ (𝓝 0 : Filter E) := - mem_of_superset hp (closedBall_mono <| le_max_right _ _) - refine' continuousAt_zero_of_forall' _ - intro ε hε - rcases exists_norm_lt 𝕜 (div_pos hε hr') with ⟨k, hk0, hkε⟩ - have hk0' := norm_pos_iff.mp hk0 - have := (set_smul_mem_nhds_zero_iff hk0').mpr hp' - refine' Filter.mem_of_superset this (smul_set_subset_iff.mpr fun x hx => _) - rw [mem_closedBall_zero, map_smul_eq_mul, ← div_mul_cancel ε hr'.ne.symm] - gcongr - exact p.mem_closedBall_zero.mp hx + refine continuousAt_zero_of_forall' fun ε hε ↦ ?_ + obtain ⟨k, hk₀, hk⟩ : ∃ k : 𝕜, 0 < ‖k‖ ∧ ‖k‖ * r < ε := by + rcases le_or_lt r 0 with hr | hr + · use 1; simpa using hr.trans_lt hε + · simpa [lt_div_iff hr] using exists_norm_lt 𝕜 (div_pos hε hr) + rw [← set_smul_mem_nhds_zero_iff (norm_pos_iff.1 hk₀), smul_closedBall_zero hk₀] at hp + exact mem_of_superset hp <| p.closedBall_mono hk.le #align seminorm.continuous_at_zero' Seminorm.continuousAt_zero' /-- A seminorm is continuous at `0` if `p.ball 0 r ∈ 𝓝 0` for *all* `r > 0`. @@ -1289,6 +1283,28 @@ lemma ball_mem_nhds [TopologicalSpace E] {p : Seminorm 𝕝 E} (hp : Continuous have this : Tendsto p (𝓝 0) (𝓝 0) := map_zero p ▸ hp.tendsto 0 by simpa only [p.ball_zero_eq] using this (Iio_mem_nhds hr) +lemma uniformSpace_eq_of_hasBasis + {ι} [UniformSpace E] [UniformAddGroup E] [ContinuousConstSMul 𝕜 E] + {p' : ι → Prop} {s : ι → Set E} (p : Seminorm 𝕜 E) (hb : (𝓝 0 : Filter E).HasBasis p' s) + (h₁ : ∃ r, p.closedBall 0 r ∈ 𝓝 0) (h₂ : ∀ i, p' i → ∃ r > 0, p.ball 0 r ⊆ s i) : + ‹UniformSpace E› = p.toAddGroupSeminorm.toSeminormedAddGroup.toUniformSpace := by + refine UniformAddGroup.ext ‹_› p.toAddGroupSeminorm.toSeminormedAddCommGroup.to_uniformAddGroup ?_ + apply le_antisymm + · rw [← @comap_norm_nhds_zero E p.toAddGroupSeminorm.toSeminormedAddGroup, ← tendsto_iff_comap] + suffices Continuous p from this.tendsto' 0 _ (map_zero p) + rcases h₁ with ⟨r, hr⟩ + exact p.continuous' hr + · rw [(@NormedAddCommGroup.nhds_zero_basis_norm_lt E + p.toAddGroupSeminorm.toSeminormedAddGroup).le_basis_iff hb] + simpa only [subset_def, mem_ball_zero] using h₂ + +lemma uniformity_eq_of_hasBasis + {ι} [UniformSpace E] [UniformAddGroup E] [ContinuousConstSMul 𝕜 E] + {p' : ι → Prop} {s : ι → Set E} (p : Seminorm 𝕜 E) (hb : (𝓝 0 : Filter E).HasBasis p' s) + (h₁ : ∃ r, p.closedBall 0 r ∈ 𝓝 0) (h₂ : ∀ i, p' i → ∃ r > 0, p.ball 0 r ⊆ s i) : + 𝓤 E = ⨅ r > 0, 𝓟 {x | p (x.1 - x.2) < r} := by + rw [uniformSpace_eq_of_hasBasis p hb h₁ h₂]; rfl + end Continuity section ShellLemmas diff --git a/Mathlib/Analysis/SpecialFunctions/Complex/Log.lean b/Mathlib/Analysis/SpecialFunctions/Complex/Log.lean index cf61729cf615c..cc12497371ab1 100644 --- a/Mathlib/Analysis/SpecialFunctions/Complex/Log.lean +++ b/Mathlib/Analysis/SpecialFunctions/Complex/Log.lean @@ -169,7 +169,7 @@ theorem exp_eq_exp_iff_exists_int {x y : ℂ} : exp x = exp y ↔ ∃ n : ℤ, x theorem countable_preimage_exp {s : Set ℂ} : (exp ⁻¹' s).Countable ↔ s.Countable := by refine' ⟨fun hs => _, fun hs => _⟩ · refine' ((hs.image exp).insert 0).mono _ - rw [Set.image_preimage_eq_inter_range, range_exp, ←Set.diff_eq, ←Set.union_singleton, + rw [Set.image_preimage_eq_inter_range, range_exp, ← Set.diff_eq, ← Set.union_singleton, Set.diff_union_self] exact Set.subset_union_left _ _ · rw [← Set.biUnion_preimage_singleton] diff --git a/Mathlib/Analysis/SpecialFunctions/Exp.lean b/Mathlib/Analysis/SpecialFunctions/Exp.lean index 4b011ac94b765..7aacf5b2df0a8 100644 --- a/Mathlib/Analysis/SpecialFunctions/Exp.lean +++ b/Mathlib/Analysis/SpecialFunctions/Exp.lean @@ -307,12 +307,12 @@ theorem comap_exp_atTop : comap exp atTop = atTop := by @[simp] theorem tendsto_exp_comp_atTop {f : α → ℝ} : Tendsto (fun x => exp (f x)) l atTop ↔ Tendsto f l atTop := by - simp_rw [←comp_apply (f := exp), ← tendsto_comap_iff, comap_exp_atTop] + simp_rw [← comp_apply (f := exp), ← tendsto_comap_iff, comap_exp_atTop] #align real.tendsto_exp_comp_at_top Real.tendsto_exp_comp_atTop theorem tendsto_comp_exp_atTop {f : ℝ → α} : Tendsto (fun x => f (exp x)) atTop l ↔ Tendsto f atTop l := by - simp_rw [←comp_apply (g := exp), ← tendsto_map'_iff, map_exp_atTop] + simp_rw [← comp_apply (g := exp), ← tendsto_map'_iff, map_exp_atTop] #align real.tendsto_comp_exp_at_top Real.tendsto_comp_exp_atTop @[simp] @@ -339,7 +339,7 @@ theorem comap_exp_nhds_zero : comap exp (𝓝 0) = atBot := @[simp] theorem tendsto_exp_comp_nhds_zero {f : α → ℝ} : Tendsto (fun x => exp (f x)) l (𝓝 0) ↔ Tendsto f l atBot := by - simp_rw [←comp_apply (f := exp), ← tendsto_comap_iff, comap_exp_nhds_zero] + simp_rw [← comp_apply (f := exp), ← tendsto_comap_iff, comap_exp_nhds_zero] #align real.tendsto_exp_comp_nhds_zero Real.tendsto_exp_comp_nhds_zero -- Porting note: new lemma @@ -443,7 +443,7 @@ theorem comap_exp_nhdsWithin_zero : comap exp (𝓝[≠] 0) = comap re atBot := theorem tendsto_exp_nhds_zero_iff {α : Type*} {l : Filter α} {f : α → ℂ} : Tendsto (fun x => exp (f x)) l (𝓝 0) ↔ Tendsto (fun x => re (f x)) l atBot := by - simp_rw [←comp_apply (f := exp), ← tendsto_comap_iff, comap_exp_nhds_zero, tendsto_comap_iff] + simp_rw [← comp_apply (f := exp), ← tendsto_comap_iff, comap_exp_nhds_zero, tendsto_comap_iff] rfl #align complex.tendsto_exp_nhds_zero_iff Complex.tendsto_exp_nhds_zero_iff diff --git a/Mathlib/Analysis/SpecialFunctions/Gaussian.lean b/Mathlib/Analysis/SpecialFunctions/Gaussian.lean index e836c2efd1058..0b2750511db20 100644 --- a/Mathlib/Analysis/SpecialFunctions/Gaussian.lean +++ b/Mathlib/Analysis/SpecialFunctions/Gaussian.lean @@ -151,7 +151,7 @@ theorem integrable_exp_neg_mul_sq {b : ℝ} (hb : 0 < b) : Integrable fun x : theorem integrableOn_Ioi_exp_neg_mul_sq_iff {b : ℝ} : IntegrableOn (fun x : ℝ => exp (-b * x ^ 2)) (Ioi 0) ↔ 0 < b := by refine' ⟨fun h => _, fun h => (integrable_exp_neg_mul_sq h).integrableOn⟩ - by_contra' hb + by_contra! hb have : ∫⁻ _ : ℝ in Ioi 0, 1 ≤ ∫⁻ x : ℝ in Ioi 0, ‖exp (-b * x ^ 2)‖₊ := by apply lintegral_mono (fun x ↦ _) simp only [neg_mul, ENNReal.one_le_coe_iff, ← toNNReal_one, toNNReal_le_iff_le_coe, @@ -172,7 +172,7 @@ theorem integrable_mul_exp_neg_mul_sq {b : ℝ} (hb : 0 < b) : theorem norm_cexp_neg_mul_sq (b : ℂ) (x : ℝ) : ‖Complex.exp (-b * (x : ℂ) ^ 2)‖ = exp (-b.re * x ^ 2) := by - rw [Complex.norm_eq_abs, Complex.abs_exp, ←ofReal_pow, mul_comm (-b) _, ofReal_mul_re, neg_re, + rw [Complex.norm_eq_abs, Complex.abs_exp, ← ofReal_pow, mul_comm (-b) _, ofReal_mul_re, neg_re, mul_comm] #align norm_cexp_neg_mul_sq norm_cexp_neg_mul_sq diff --git a/Mathlib/Analysis/SpecialFunctions/Log/Base.lean b/Mathlib/Analysis/SpecialFunctions/Log/Base.lean index 870fa88ec0597..dfb6e645e328d 100644 --- a/Mathlib/Analysis/SpecialFunctions/Log/Base.lean +++ b/Mathlib/Analysis/SpecialFunctions/Log/Base.lean @@ -454,7 +454,7 @@ lemma Real.induction_Ico_mul {P : ℝ → Prop} (x₀ r : ℝ) (hr : 1 < r) (hx · intro x hx have hx' : 0 < x / x₀ := div_pos (hx₀.trans_le hx) hx₀ refine this ⌊logb r (x / x₀)⌋₊ x ?_ - rw [mem_Ico, ←div_lt_iff hx₀, ← rpow_nat_cast, ←logb_lt_iff_lt_rpow hr hx', Nat.cast_add, + rw [mem_Ico, ← div_lt_iff hx₀, ← rpow_nat_cast, ← logb_lt_iff_lt_rpow hr hx', Nat.cast_add, Nat.cast_one] exact ⟨hx, Nat.lt_floor_add_one _⟩ intro n diff --git a/Mathlib/Analysis/SpecialFunctions/Log/Basic.lean b/Mathlib/Analysis/SpecialFunctions/Log/Basic.lean index 9dd7b39c8f877..487700b077ee6 100644 --- a/Mathlib/Analysis/SpecialFunctions/Log/Basic.lean +++ b/Mathlib/Analysis/SpecialFunctions/Log/Basic.lean @@ -184,7 +184,7 @@ theorem log_pos (hx : 1 < x) : 0 < log x := #align real.log_pos Real.log_pos theorem log_pos_of_lt_neg_one (hx : x < -1) : 0 < log x := by - rw [←neg_neg x, log_neg_eq_log] + rw [← neg_neg x, log_neg_eq_log] have : 1 < -x := by linarith exact log_pos this @@ -198,7 +198,7 @@ theorem log_neg (h0 : 0 < x) (h1 : x < 1) : log x < 0 := #align real.log_neg Real.log_neg theorem log_neg_of_lt_zero (h0 : x < 0) (h1 : -1 < x) : log x < 0 := by - rw [←neg_neg x, log_neg_eq_log] + rw [← neg_neg x, log_neg_eq_log] have h0' : 0 < -x := by linarith have h1' : -x < 1 := by linarith exact log_neg h0' h1' @@ -231,7 +231,7 @@ theorem log_nat_cast_nonneg (n : ℕ) : 0 ≤ log n := by exact log_nonneg this theorem log_neg_nat_cast_nonneg (n : ℕ) : 0 ≤ log (-n) := by - rw [←log_neg_eq_log, neg_neg] + rw [← log_neg_eq_log, neg_neg] exact log_nat_cast_nonneg _ theorem log_int_cast_nonneg (n : ℤ) : 0 ≤ log n := by @@ -243,8 +243,8 @@ theorem log_int_cast_nonneg (n : ℤ) : 0 ≤ log n := by cases hn with | inl hn => simp [hn.symm] | inr hn => - have : (1 : ℝ) ≤ -n := by rw [←neg_zero, ←lt_neg] at hn; exact mod_cast hn - rw [←log_neg_eq_log] + have : (1 : ℝ) ≤ -n := by rw [← neg_zero, ← lt_neg] at hn; exact mod_cast hn + rw [← log_neg_eq_log] exact log_nonneg this theorem strictMonoOn_log : StrictMonoOn log (Set.Ioi 0) := fun _ hx _ _ hxy => log_lt_log hx hxy @@ -511,21 +511,21 @@ lemma log_pos_of_isNegNat (h : NormNum.IsInt e (.negOfNat n)) (w : Nat.blt 1 n = lemma log_pos_of_isRat : (NormNum.IsRat e n d) → (decide ((1 : ℚ) < n / d)) → (0 < Real.log (e : ℝ)) | ⟨inv, eq⟩, h => by - rw [eq, invOf_eq_inv, ←div_eq_mul_inv] + rw [eq, invOf_eq_inv, ← div_eq_mul_inv] have : 1 < (n : ℝ) / d := by exact_mod_cast of_decide_eq_true h exact Real.log_pos this lemma log_pos_of_isRat_neg : (NormNum.IsRat e n d) → (decide (n / d < (-1 : ℚ))) → (0 < Real.log (e : ℝ)) | ⟨inv, eq⟩, h => by - rw [eq, invOf_eq_inv, ←div_eq_mul_inv] + rw [eq, invOf_eq_inv, ← div_eq_mul_inv] have : (n : ℝ) / d < -1 := by exact_mod_cast of_decide_eq_true h exact Real.log_pos_of_lt_neg_one this lemma log_nz_of_isRat : (NormNum.IsRat e n d) → (decide ((0 : ℚ) < n / d)) → (decide (n / d < (1 : ℚ))) → (Real.log (e : ℝ) ≠ 0) | ⟨inv, eq⟩, h₁, h₂ => by - rw [eq, invOf_eq_inv, ←div_eq_mul_inv] + rw [eq, invOf_eq_inv, ← div_eq_mul_inv] have h₁' : 0 < (n : ℝ) / d := by exact_mod_cast of_decide_eq_true h₁ have h₂' : (n : ℝ) / d < 1 := by exact_mod_cast of_decide_eq_true h₂ exact ne_of_lt <| Real.log_neg h₁' h₂' @@ -533,7 +533,7 @@ lemma log_nz_of_isRat : (NormNum.IsRat e n d) → (decide ((0 : ℚ) < n / d)) lemma log_nz_of_isRat_neg : (NormNum.IsRat e n d) → (decide (n / d < (0 : ℚ))) → (decide ((-1 : ℚ) < n / d)) → (Real.log (e : ℝ) ≠ 0) | ⟨inv, eq⟩, h₁, h₂ => by - rw [eq, invOf_eq_inv, ←div_eq_mul_inv] + rw [eq, invOf_eq_inv, ← div_eq_mul_inv] have h₁' : (n : ℝ) / d < 0 := by exact_mod_cast of_decide_eq_true h₁ have h₂' : -1 < (n : ℝ) / d := by exact_mod_cast of_decide_eq_true h₂ exact ne_of_lt <| Real.log_neg_of_lt_zero h₁' h₂' @@ -557,7 +557,7 @@ def evalLogIntCast : PositivityExt where eval {_ _} _zα _pα e := do def evalLogNatLit : PositivityExt where eval {_ _} _zα _pα e := do let .app (f : Q(ℝ → ℝ)) (a : Q(ℝ)) ← withReducible (whnf e) | throwError "not Real.log" guard <|← withDefault <| withNewMCtxDepth <| isDefEq f q(Real.log) - match ←NormNum.derive a with + match ← NormNum.derive a with | .isNat (_ : Q(AddMonoidWithOne ℝ)) lit p => assumeInstancesCommute have p : Q(NormNum.IsNat $a $lit) := p diff --git a/Mathlib/Analysis/SpecialFunctions/Log/Deriv.lean b/Mathlib/Analysis/SpecialFunctions/Log/Deriv.lean index c823133f5072a..1f2a89ac8d0eb 100644 --- a/Mathlib/Analysis/SpecialFunctions/Log/Deriv.lean +++ b/Mathlib/Analysis/SpecialFunctions/Log/Deriv.lean @@ -77,7 +77,7 @@ theorem deriv_log' : deriv log = Inv.inv := theorem contDiffOn_log {n : ℕ∞} : ContDiffOn ℝ n log {0}ᶜ := by suffices : ContDiffOn ℝ ⊤ log {0}ᶜ; exact this.of_le le_top - refine' (contDiffOn_top_iff_deriv_of_open isOpen_compl_singleton).2 _ + refine' (contDiffOn_top_iff_deriv_of_isOpen isOpen_compl_singleton).2 _ simp [differentiableOn_log, contDiffOn_inv] #align real.cont_diff_on_log Real.contDiffOn_log diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Asymptotics.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Asymptotics.lean index 1c14f3a142698..39f42c5b2e481 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Asymptotics.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Asymptotics.lean @@ -58,12 +58,12 @@ lemma tendsto_rpow_atTop_of_base_lt_one (b : ℝ) (hb₀ : -1 < b) (hb₁ : b < rcases lt_trichotomy b 0 with hb|rfl|hb case inl => -- b < 0 simp_rw [Real.rpow_def_of_nonpos hb.le, hb.ne, ite_false] - rw [←isLittleO_const_iff (c := (1:ℝ)) one_ne_zero, (one_mul (1 : ℝ)).symm] + rw [← isLittleO_const_iff (c := (1:ℝ)) one_ne_zero, (one_mul (1 : ℝ)).symm] refine IsLittleO.mul_isBigO ?exp ?cos case exp => rw [isLittleO_const_iff one_ne_zero] refine tendsto_exp_atBot.comp <| (tendsto_const_mul_atBot_of_neg ?_).mpr tendsto_id - rw [←log_neg_eq_log, log_neg_iff (by linarith)] + rw [← log_neg_eq_log, log_neg_iff (by linarith)] linarith case cos => rw [isBigO_iff] diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean b/Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean index 737848b381fbb..920e9ce3a48f8 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean @@ -152,7 +152,7 @@ theorem list_prod_map_rpow (l : List ℝ≥0) (r : ℝ) : theorem list_prod_map_rpow' {ι} (l : List ι) (f : ι → ℝ≥0) (r : ℝ) : (l.map (f · ^ r)).prod = (l.map f).prod ^ r := by - rw [←list_prod_map_rpow, List.map_map]; rfl + rw [← list_prod_map_rpow, List.map_map]; rfl /-- `rpow` version of `Multiset.prod_map_pow` for `ℝ≥0`. -/ lemma multiset_prod_map_rpow {ι} (s : Multiset ι) (f : ι → ℝ≥0) (r : ℝ) : @@ -180,7 +180,7 @@ theorem _root_.Real.list_prod_map_rpow (l : List ℝ) (hl : ∀ x ∈ l, (0 : theorem _root_.Real.list_prod_map_rpow' {ι} (l : List ι) (f : ι → ℝ) (hl : ∀ i ∈ l, (0 : ℝ) ≤ f i) (r : ℝ) : (l.map (f · ^ r)).prod = (l.map f).prod ^ r := by - rw [←Real.list_prod_map_rpow (l.map f) _ r, List.map_map]; rfl + rw [← Real.list_prod_map_rpow (l.map f) _ r, List.map_map]; rfl simpa using hl /-- `rpow` version of `Multiset.prod_map_pow`. -/ diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean index ca3470ba45460..3e2883de10cec 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean @@ -506,7 +506,7 @@ theorem rpow_le_rpow_of_exponent_le (hx : 1 ≤ x) (hyz : y ≤ z) : x ^ y ≤ x theorem rpow_lt_rpow_of_exponent_neg {x y z : ℝ} (hy : 0 < y) (hxy : y < x) (hz : z < 0) : x ^ z < y ^ z := by have hx : 0 < x := hy.trans hxy - rw [←neg_neg z, Real.rpow_neg (le_of_lt hx) (-z), Real.rpow_neg (le_of_lt hy) (-z), + rw [← neg_neg z, Real.rpow_neg (le_of_lt hx) (-z), Real.rpow_neg (le_of_lt hy) (-z), inv_lt_inv (rpow_pos_of_pos hx _) (rpow_pos_of_pos hy _)] exact Real.rpow_lt_rpow (by positivity) hxy <| neg_pos_of_neg hz diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Angle.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Angle.lean index c9b07ab00a9f3..004101c3955f2 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Angle.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Angle.lean @@ -229,7 +229,7 @@ theorem two_nsmul_eq_pi_iff {θ : Angle} : (2 : ℕ) • θ = π ↔ θ = (π / rw [coe_nsmul, two_nsmul_eq_iff] -- Porting note: `congr` didn't simplify the goal of iff of `Or`s convert Iff.rfl - rw [add_comm, ← coe_add, ← sub_eq_zero, ← coe_sub, neg_div, ←neg_sub, sub_neg_eq_add, add_assoc, + rw [add_comm, ← coe_add, ← sub_eq_zero, ← coe_sub, neg_div, ← neg_sub, sub_neg_eq_add, add_assoc, add_halves, ← two_mul, coe_neg, coe_two_pi, neg_zero] #align real.angle.two_nsmul_eq_pi_iff Real.Angle.two_nsmul_eq_pi_iff diff --git a/Mathlib/Analysis/SpecificLimits/Basic.lean b/Mathlib/Analysis/SpecificLimits/Basic.lean index 1f01fe9675837..97b565219fce2 100644 --- a/Mathlib/Analysis/SpecificLimits/Basic.lean +++ b/Mathlib/Analysis/SpecificLimits/Basic.lean @@ -28,40 +28,43 @@ open Classical Topology Nat BigOperators uniformity NNReal ENNReal variable {α : Type*} {β : Type*} {ι : Type*} -theorem tendsto_inverse_atTop_nhds_0_nat : Tendsto (fun n : ℕ => (n : ℝ)⁻¹) atTop (𝓝 0) := +theorem tendsto_inverse_atTop_nhds_0_nat : Tendsto (fun n : ℕ ↦ (n : ℝ)⁻¹) atTop (𝓝 0) := tendsto_inv_atTop_zero.comp tendsto_nat_cast_atTop_atTop #align tendsto_inverse_at_top_nhds_0_nat tendsto_inverse_atTop_nhds_0_nat -theorem tendsto_const_div_atTop_nhds_0_nat (C : ℝ) : Tendsto (fun n : ℕ => C / n) atTop (𝓝 0) := by +theorem tendsto_const_div_atTop_nhds_0_nat (C : ℝ) : Tendsto (fun n : ℕ ↦ C / n) atTop (𝓝 0) := by simpa only [mul_zero] using tendsto_const_nhds.mul tendsto_inverse_atTop_nhds_0_nat #align tendsto_const_div_at_top_nhds_0_nat tendsto_const_div_atTop_nhds_0_nat +theorem tendsto_one_div_atTop_nhds_0_nat : Tendsto (fun n : ℕ ↦ 1/(n : ℝ)) atTop (𝓝 0) := + tendsto_const_div_atTop_nhds_0_nat 1 + theorem NNReal.tendsto_inverse_atTop_nhds_0_nat : - Tendsto (fun n : ℕ => (n : ℝ≥0)⁻¹) atTop (𝓝 0) := by + Tendsto (fun n : ℕ ↦ (n : ℝ≥0)⁻¹) atTop (𝓝 0) := by rw [← NNReal.tendsto_coe] exact _root_.tendsto_inverse_atTop_nhds_0_nat #align nnreal.tendsto_inverse_at_top_nhds_0_nat NNReal.tendsto_inverse_atTop_nhds_0_nat theorem NNReal.tendsto_const_div_atTop_nhds_0_nat (C : ℝ≥0) : - Tendsto (fun n : ℕ => C / n) atTop (𝓝 0) := by + Tendsto (fun n : ℕ ↦ C / n) atTop (𝓝 0) := by simpa using tendsto_const_nhds.mul NNReal.tendsto_inverse_atTop_nhds_0_nat #align nnreal.tendsto_const_div_at_top_nhds_0_nat NNReal.tendsto_const_div_atTop_nhds_0_nat theorem tendsto_one_div_add_atTop_nhds_0_nat : - Tendsto (fun n : ℕ => 1 / ((n : ℝ) + 1)) atTop (𝓝 0) := - suffices Tendsto (fun n : ℕ => 1 / (↑(n + 1) : ℝ)) atTop (𝓝 0) by simpa + Tendsto (fun n : ℕ ↦ 1 / ((n : ℝ) + 1)) atTop (𝓝 0) := + suffices Tendsto (fun n : ℕ ↦ 1 / (↑(n + 1) : ℝ)) atTop (𝓝 0) by simpa (tendsto_add_atTop_iff_nat 1).2 (_root_.tendsto_const_div_atTop_nhds_0_nat 1) #align tendsto_one_div_add_at_top_nhds_0_nat tendsto_one_div_add_atTop_nhds_0_nat theorem NNReal.tendsto_algebraMap_inverse_atTop_nhds_0_nat (𝕜 : Type*) [Semiring 𝕜] [Algebra ℝ≥0 𝕜] [TopologicalSpace 𝕜] [TopologicalSemiring 𝕜] [ContinuousSMul ℝ≥0 𝕜] : - Tendsto (algebraMap ℝ≥0 𝕜 ∘ fun n : ℕ => (n : ℝ≥0)⁻¹) atTop (nhds 0) := by + Tendsto (algebraMap ℝ≥0 𝕜 ∘ fun n : ℕ ↦ (n : ℝ≥0)⁻¹) atTop (nhds 0) := by convert (continuous_algebraMap ℝ≥0 𝕜).continuousAt.tendsto.comp tendsto_inverse_atTop_nhds_0_nat rw [map_zero] theorem tendsto_algebraMap_inverse_atTop_nhds_0_nat (𝕜 : Type*) [Semiring 𝕜] [Algebra ℝ 𝕜] [TopologicalSpace 𝕜] [TopologicalSemiring 𝕜] [ContinuousSMul ℝ 𝕜] : - Tendsto (algebraMap ℝ 𝕜 ∘ fun n : ℕ => (n : ℝ)⁻¹) atTop (nhds 0) := + Tendsto (algebraMap ℝ 𝕜 ∘ fun n : ℕ ↦ (n : ℝ)⁻¹) atTop (nhds 0) := NNReal.tendsto_algebraMap_inverse_atTop_nhds_0_nat 𝕜 /-- The limit of `n / (n + x)` is 1, for any constant `x` (valid in `ℝ` or any topological division @@ -71,9 +74,9 @@ TODO: introduce a typeclass saying that `1 / n` tends to 0 at top, making it pos statement simultaneously on `ℚ`, `ℝ` and `ℂ`. -/ theorem tendsto_coe_nat_div_add_atTop {𝕜 : Type*} [DivisionRing 𝕜] [TopologicalSpace 𝕜] [CharZero 𝕜] [Algebra ℝ 𝕜] [ContinuousSMul ℝ 𝕜] [TopologicalDivisionRing 𝕜] (x : 𝕜) : - Tendsto (fun n : ℕ => (n : 𝕜) / (n + x)) atTop (𝓝 1) := by - refine' Tendsto.congr' ((eventually_ne_atTop 0).mp (eventually_of_forall fun n hn => _)) _ - · exact fun n : ℕ => 1 / (1 + x / n) + Tendsto (fun n : ℕ ↦ (n : 𝕜) / (n + x)) atTop (𝓝 1) := by + refine' Tendsto.congr' ((eventually_ne_atTop 0).mp (eventually_of_forall fun n hn ↦ _)) _ + · exact fun n : ℕ ↦ 1 / (1 + x / n) · field_simp [Nat.cast_ne_zero.mpr hn] · have : 𝓝 (1 : 𝕜) = 𝓝 (1 / (1 + x * (0 : 𝕜))) := by rw [mul_zero, add_zero, div_one] @@ -92,35 +95,35 @@ theorem tendsto_coe_nat_div_add_atTop {𝕜 : Type*} [DivisionRing 𝕜] [Topolo theorem tendsto_add_one_pow_atTop_atTop_of_pos [LinearOrderedSemiring α] [Archimedean α] {r : α} - (h : 0 < r) : Tendsto (fun n : ℕ => (r + 1) ^ n) atTop atTop := - (tendsto_atTop_atTop_of_monotone' fun _ _ => pow_le_pow (le_add_of_nonneg_left (le_of_lt h))) <| - not_bddAbove_iff.2 fun _ => Set.exists_range_iff.2 <| add_one_pow_unbounded_of_pos _ h + (h : 0 < r) : Tendsto (fun n : ℕ ↦ (r + 1) ^ n) atTop atTop := + (tendsto_atTop_atTop_of_monotone' fun _ _ ↦ pow_le_pow (le_add_of_nonneg_left (le_of_lt h))) <| + not_bddAbove_iff.2 fun _ ↦ Set.exists_range_iff.2 <| add_one_pow_unbounded_of_pos _ h #align tendsto_add_one_pow_at_top_at_top_of_pos tendsto_add_one_pow_atTop_atTop_of_pos theorem tendsto_pow_atTop_atTop_of_one_lt [LinearOrderedRing α] [Archimedean α] {r : α} - (h : 1 < r) : Tendsto (fun n : ℕ => r ^ n) atTop atTop := + (h : 1 < r) : Tendsto (fun n : ℕ ↦ r ^ n) atTop atTop := sub_add_cancel r 1 ▸ tendsto_add_one_pow_atTop_atTop_of_pos (sub_pos.2 h) #align tendsto_pow_at_top_at_top_of_one_lt tendsto_pow_atTop_atTop_of_one_lt theorem Nat.tendsto_pow_atTop_atTop_of_one_lt {m : ℕ} (h : 1 < m) : - Tendsto (fun n : ℕ => m ^ n) atTop atTop := + Tendsto (fun n : ℕ ↦ m ^ n) atTop atTop := tsub_add_cancel_of_le (le_of_lt h) ▸ tendsto_add_one_pow_atTop_atTop_of_pos (tsub_pos_of_lt h) #align nat.tendsto_pow_at_top_at_top_of_one_lt Nat.tendsto_pow_atTop_atTop_of_one_lt theorem tendsto_pow_atTop_nhds_0_of_lt_1 {𝕜 : Type*} [LinearOrderedField 𝕜] [Archimedean 𝕜] [TopologicalSpace 𝕜] [OrderTopology 𝕜] {r : 𝕜} (h₁ : 0 ≤ r) (h₂ : r < 1) : - Tendsto (fun n : ℕ => r ^ n) atTop (𝓝 0) := + Tendsto (fun n : ℕ ↦ r ^ n) atTop (𝓝 0) := h₁.eq_or_lt.elim - (fun hr => (tendsto_add_atTop_iff_nat 1).mp <| by + (fun hr ↦ (tendsto_add_atTop_iff_nat 1).mp <| by simp [_root_.pow_succ, ← hr, tendsto_const_nhds]) - (fun hr => + (fun hr ↦ have := one_lt_inv hr h₂ |> tendsto_pow_atTop_atTop_of_one_lt - (tendsto_inv_atTop_zero.comp this).congr fun n => by simp) + (tendsto_inv_atTop_zero.comp this).congr fun n ↦ by simp) #align tendsto_pow_at_top_nhds_0_of_lt_1 tendsto_pow_atTop_nhds_0_of_lt_1 @[simp] theorem tendsto_pow_atTop_nhds_0_iff {𝕜 : Type*} [LinearOrderedField 𝕜] [Archimedean 𝕜] [TopologicalSpace 𝕜] [OrderTopology 𝕜] {r : 𝕜} : - Tendsto (fun n : ℕ => r ^ n) atTop (𝓝 0) ↔ |r| < 1 := by + Tendsto (fun n : ℕ ↦ r ^ n) atTop (𝓝 0) ↔ |r| < 1 := by rw [tendsto_zero_iff_abs_tendsto_zero] refine ⟨fun h ↦ by_contra (fun hr_le ↦ ?_), fun h ↦ ?_⟩ · by_cases hr : 1 = |r| @@ -136,17 +139,17 @@ theorem tendsto_pow_atTop_nhds_0_of_lt_1 {𝕜 : Type*} [LinearOrderedField 𝕜 theorem tendsto_pow_atTop_nhdsWithin_0_of_lt_1 {𝕜 : Type*} [LinearOrderedField 𝕜] [Archimedean 𝕜] [TopologicalSpace 𝕜] [OrderTopology 𝕜] {r : 𝕜} (h₁ : 0 < r) (h₂ : r < 1) : - Tendsto (fun n : ℕ => r ^ n) atTop (𝓝[>] 0) := + Tendsto (fun n : ℕ ↦ r ^ n) atTop (𝓝[>] 0) := tendsto_inf.2 ⟨tendsto_pow_atTop_nhds_0_of_lt_1 h₁.le h₂, - tendsto_principal.2 <| eventually_of_forall fun _ => pow_pos h₁ _⟩ + tendsto_principal.2 <| eventually_of_forall fun _ ↦ pow_pos h₁ _⟩ #align tendsto_pow_at_top_nhds_within_0_of_lt_1 tendsto_pow_atTop_nhdsWithin_0_of_lt_1 theorem uniformity_basis_dist_pow_of_lt_1 {α : Type*} [PseudoMetricSpace α] {r : ℝ} (h₀ : 0 < r) (h₁ : r < 1) : - (uniformity α).HasBasis (fun _ : ℕ => True) fun k => { p : α × α | dist p.1 p.2 < r ^ k } := - Metric.mk_uniformity_basis (fun _ _ => pow_pos h₀ _) fun _ ε0 => - (exists_pow_lt_of_lt_one ε0 h₁).imp fun _ hk => ⟨trivial, hk.le⟩ + (uniformity α).HasBasis (fun _ : ℕ ↦ True) fun k ↦ { p : α × α | dist p.1 p.2 < r ^ k } := + Metric.mk_uniformity_basis (fun _ _ ↦ pow_pos h₀ _) fun _ ε0 ↦ + (exists_pow_lt_of_lt_one ε0 h₁).imp fun _ hk ↦ ⟨trivial, hk.le⟩ #align uniformity_basis_dist_pow_of_lt_1 uniformity_basis_dist_pow_of_lt_1 theorem geom_lt {u : ℕ → ℝ} {c : ℝ} (hc : 0 ≤ c) {n : ℕ} (hn : 0 < n) @@ -179,19 +182,19 @@ theorem le_geom {u : ℕ → ℝ} {c : ℝ} (hc : 0 ≤ c) (n : ℕ) (h : ∀ k then it goes to +∞. -/ theorem tendsto_atTop_of_geom_le {v : ℕ → ℝ} {c : ℝ} (h₀ : 0 < v 0) (hc : 1 < c) (hu : ∀ n, c * v n ≤ v (n + 1)) : Tendsto v atTop atTop := - (tendsto_atTop_mono fun n => geom_le (zero_le_one.trans hc.le) n fun k _ => hu k) <| + (tendsto_atTop_mono fun n ↦ geom_le (zero_le_one.trans hc.le) n fun k _ ↦ hu k) <| (tendsto_pow_atTop_atTop_of_one_lt hc).atTop_mul_const h₀ #align tendsto_at_top_of_geom_le tendsto_atTop_of_geom_le theorem NNReal.tendsto_pow_atTop_nhds_0_of_lt_1 {r : ℝ≥0} (hr : r < 1) : - Tendsto (fun n : ℕ => r ^ n) atTop (𝓝 0) := + Tendsto (fun n : ℕ ↦ r ^ n) atTop (𝓝 0) := NNReal.tendsto_coe.1 <| by simp only [NNReal.coe_pow, NNReal.coe_zero, _root_.tendsto_pow_atTop_nhds_0_of_lt_1 r.coe_nonneg hr] #align nnreal.tendsto_pow_at_top_nhds_0_of_lt_1 NNReal.tendsto_pow_atTop_nhds_0_of_lt_1 theorem ENNReal.tendsto_pow_atTop_nhds_0_of_lt_1 {r : ℝ≥0∞} (hr : r < 1) : - Tendsto (fun n : ℕ => r ^ n) atTop (𝓝 0) := by + Tendsto (fun n : ℕ ↦ r ^ n) atTop (𝓝 0) := by rcases ENNReal.lt_iff_exists_coe.1 hr with ⟨r, rfl, hr'⟩ rw [← ENNReal.coe_zero] norm_cast at * @@ -204,16 +207,16 @@ theorem ENNReal.tendsto_pow_atTop_nhds_0_of_lt_1 {r : ℝ≥0∞} (hr : r < 1) : section Geometric theorem hasSum_geometric_of_lt_1 {r : ℝ} (h₁ : 0 ≤ r) (h₂ : r < 1) : - HasSum (fun n : ℕ => r ^ n) (1 - r)⁻¹ := + HasSum (fun n : ℕ ↦ r ^ n) (1 - r)⁻¹ := have : r ≠ 1 := ne_of_lt h₂ - have : Tendsto (fun n => (r ^ n - 1) * (r - 1)⁻¹) atTop (𝓝 ((0 - 1) * (r - 1)⁻¹)) := + have : Tendsto (fun n ↦ (r ^ n - 1) * (r - 1)⁻¹) atTop (𝓝 ((0 - 1) * (r - 1)⁻¹)) := ((tendsto_pow_atTop_nhds_0_of_lt_1 h₁ h₂).sub tendsto_const_nhds).mul tendsto_const_nhds (hasSum_iff_tendsto_nat_of_nonneg (pow_nonneg h₁) _).mpr <| by simp_all [neg_inv, geom_sum_eq, div_eq_mul_inv] #align has_sum_geometric_of_lt_1 hasSum_geometric_of_lt_1 theorem summable_geometric_of_lt_1 {r : ℝ} (h₁ : 0 ≤ r) (h₂ : r < 1) : - Summable fun n : ℕ => r ^ n := + Summable fun n : ℕ ↦ r ^ n := ⟨_, hasSum_geometric_of_lt_1 h₁ h₂⟩ #align summable_geometric_of_lt_1 summable_geometric_of_lt_1 @@ -221,16 +224,16 @@ theorem tsum_geometric_of_lt_1 {r : ℝ} (h₁ : 0 ≤ r) (h₂ : r < 1) : ∑' (hasSum_geometric_of_lt_1 h₁ h₂).tsum_eq #align tsum_geometric_of_lt_1 tsum_geometric_of_lt_1 -theorem hasSum_geometric_two : HasSum (fun n : ℕ => ((1 : ℝ) / 2) ^ n) 2 := by +theorem hasSum_geometric_two : HasSum (fun n : ℕ ↦ ((1 : ℝ) / 2) ^ n) 2 := by convert hasSum_geometric_of_lt_1 _ _ <;> norm_num #align has_sum_geometric_two hasSum_geometric_two -theorem summable_geometric_two : Summable fun n : ℕ => ((1 : ℝ) / 2) ^ n := +theorem summable_geometric_two : Summable fun n : ℕ ↦ ((1 : ℝ) / 2) ^ n := ⟨_, hasSum_geometric_two⟩ #align summable_geometric_two summable_geometric_two theorem summable_geometric_two_encode {ι : Type*} [Encodable ι] : - Summable fun i : ι => (1 / 2 : ℝ) ^ Encodable.encode i := + Summable fun i : ι ↦ (1 / 2 : ℝ) ^ Encodable.encode i := summable_geometric_two.comp_injective Encodable.encode_injective #align summable_geometric_two_encode summable_geometric_two_encode @@ -243,7 +246,7 @@ theorem sum_geometric_two_le (n : ℕ) : (∑ i : ℕ in range n, (1 / (2 : ℝ) intro i apply pow_nonneg norm_num - convert sum_le_tsum (range n) (fun i _ => this i) summable_geometric_two + convert sum_le_tsum (range n) (fun i _ ↦ this i) summable_geometric_two exact tsum_geometric_two.symm #align sum_geometric_two_le sum_geometric_two_le @@ -254,17 +257,17 @@ theorem tsum_geometric_inv_two : (∑' n : ℕ, (2 : ℝ)⁻¹ ^ n) = 2 := /-- The sum of `2⁻¹ ^ i` for `n ≤ i` equals `2 * 2⁻¹ ^ n`. -/ theorem tsum_geometric_inv_two_ge (n : ℕ) : (∑' i, ite (n ≤ i) ((2 : ℝ)⁻¹ ^ i) 0) = 2 * 2⁻¹ ^ n := by - have A : Summable fun i : ℕ => ite (n ≤ i) ((2⁻¹ : ℝ) ^ i) 0 := by + have A : Summable fun i : ℕ ↦ ite (n ≤ i) ((2⁻¹ : ℝ) ^ i) 0 := by simpa only [← piecewise_eq_indicator, one_div] using summable_geometric_two.indicator {i | n ≤ i} - have B : ((Finset.range n).sum fun i : ℕ => ite (n ≤ i) ((2⁻¹ : ℝ) ^ i) 0) = 0 := - Finset.sum_eq_zero fun i hi => - ite_eq_right_iff.2 fun h => (lt_irrefl _ ((Finset.mem_range.1 hi).trans_le h)).elim + have B : ((Finset.range n).sum fun i : ℕ ↦ ite (n ≤ i) ((2⁻¹ : ℝ) ^ i) 0) = 0 := + Finset.sum_eq_zero fun i hi ↦ + ite_eq_right_iff.2 fun h ↦ (lt_irrefl _ ((Finset.mem_range.1 hi).trans_le h)).elim simp only [← _root_.sum_add_tsum_nat_add n A, B, if_true, zero_add, zero_le', le_add_iff_nonneg_left, pow_add, _root_.tsum_mul_right, tsum_geometric_inv_two] #align tsum_geometric_inv_two_ge tsum_geometric_inv_two_ge -theorem hasSum_geometric_two' (a : ℝ) : HasSum (fun n : ℕ => a / 2 / 2 ^ n) a := by +theorem hasSum_geometric_two' (a : ℝ) : HasSum (fun n : ℕ ↦ a / 2 / 2 ^ n) a := by convert HasSum.mul_left (a / 2) (hasSum_geometric_of_lt_1 (le_of_lt one_half_pos) one_half_lt_one) using 1 · funext n @@ -273,7 +276,7 @@ theorem hasSum_geometric_two' (a : ℝ) : HasSum (fun n : ℕ => a / 2 / 2 ^ n) · norm_num #align has_sum_geometric_two' hasSum_geometric_two' -theorem summable_geometric_two' (a : ℝ) : Summable fun n : ℕ => a / 2 / 2 ^ n := +theorem summable_geometric_two' (a : ℝ) : Summable fun n : ℕ ↦ a / 2 / 2 ^ n := ⟨a, hasSum_geometric_two' a⟩ #align summable_geometric_two' summable_geometric_two' @@ -282,14 +285,14 @@ theorem tsum_geometric_two' (a : ℝ) : ∑' n : ℕ, a / 2 / 2 ^ n = a := #align tsum_geometric_two' tsum_geometric_two' /-- **Sum of a Geometric Series** -/ -theorem NNReal.hasSum_geometric {r : ℝ≥0} (hr : r < 1) : HasSum (fun n : ℕ => r ^ n) (1 - r)⁻¹ := by +theorem NNReal.hasSum_geometric {r : ℝ≥0} (hr : r < 1) : HasSum (fun n : ℕ ↦ r ^ n) (1 - r)⁻¹ := by apply NNReal.hasSum_coe.1 push_cast rw [NNReal.coe_sub (le_of_lt hr)] exact hasSum_geometric_of_lt_1 r.coe_nonneg hr #align nnreal.has_sum_geometric NNReal.hasSum_geometric -theorem NNReal.summable_geometric {r : ℝ≥0} (hr : r < 1) : Summable fun n : ℕ => r ^ n := +theorem NNReal.summable_geometric {r : ℝ≥0} (hr : r < 1) : Summable fun n : ℕ ↦ r ^ n := ⟨_, NNReal.hasSum_geometric hr⟩ #align nnreal.summable_geometric NNReal.summable_geometric @@ -307,8 +310,8 @@ theorem ENNReal.tsum_geometric (r : ℝ≥0∞) : ∑' n : ℕ, r ^ n = (1 - r) convert ENNReal.tsum_coe_eq (NNReal.hasSum_geometric hr) rw [ENNReal.coe_inv <| ne_of_gt <| tsub_pos_iff_lt.2 hr, coe_sub, coe_one] · rw [tsub_eq_zero_iff_le.mpr hr, ENNReal.inv_zero, ENNReal.tsum_eq_iSup_nat, iSup_eq_top] - refine' fun a ha => - (ENNReal.exists_nat_gt (lt_top_iff_ne_top.1 ha)).imp fun n hn => lt_of_lt_of_le hn _ + refine' fun a ha ↦ + (ENNReal.exists_nat_gt (lt_top_iff_ne_top.1 ha)).imp fun n hn ↦ lt_of_lt_of_le hn _ calc (n : ℝ≥0∞) = ∑ i in range n, 1 := by rw [sum_const, nsmul_one, card_range] _ ≤ ∑ i in range n, r ^ i := by gcongr; apply one_le_pow_of_one_le' hr @@ -392,8 +395,8 @@ section LeGeometric variable [PseudoMetricSpace α] {r C : ℝ} (hr : r < 1) {f : ℕ → α} (hu : ∀ n, dist (f n) (f (n + 1)) ≤ C * r ^ n) -theorem aux_hasSum_of_le_geometric : HasSum (fun n : ℕ => C * r ^ n) (C / (1 - r)) := by - rcases sign_cases_of_C_mul_pow_nonneg fun n => dist_nonneg.trans (hu n) with (rfl | ⟨_, r₀⟩) +theorem aux_hasSum_of_le_geometric : HasSum (fun n : ℕ ↦ C * r ^ n) (C / (1 - r)) := by + rcases sign_cases_of_C_mul_pow_nonneg fun n ↦ dist_nonneg.trans (hu n) with (rfl | ⟨_, r₀⟩) · simp [hasSum_zero] · refine' HasSum.mul_left C _ simpa using hasSum_geometric_of_lt_1 r₀ hr @@ -457,8 +460,8 @@ end LeGeometric /-- A series whose terms are bounded by the terms of a converging geometric series converges. -/ theorem summable_one_div_pow_of_le {m : ℝ} {f : ℕ → ℕ} (hm : 1 < m) (fi : ∀ i, i ≤ f i) : - Summable fun i => 1 / m ^ f i := by - refine .of_nonneg_of_le (fun a => by positivity) (fun a => ?_) + Summable fun i ↦ 1 / m ^ f i := by + refine .of_nonneg_of_le (fun a ↦ by positivity) (fun a ↦ ?_) (summable_geometric_of_lt_1 (one_div_nonneg.mpr (zero_le_one.trans hm.le)) ((one_div_lt (zero_lt_one.trans hm) zero_lt_one).mpr (one_div_one.le.trans_lt hm))) rw [div_pow, one_pow] @@ -474,8 +477,8 @@ def posSumOfEncodable {ε : ℝ} (hε : 0 < ε) (ι) [Encodable ι] : { ε' : ι → ℝ // (∀ i, 0 < ε' i) ∧ ∃ c, HasSum ε' c ∧ c ≤ ε } := by let f n := ε / 2 / 2 ^ n have hf : HasSum f ε := hasSum_geometric_two' _ - have f0 : ∀ n, 0 < f n := fun n => div_pos (half_pos hε) (pow_pos zero_lt_two _) - refine' ⟨f ∘ Encodable.encode, fun i => f0 _, _⟩ + have f0 : ∀ n, 0 < f n := fun n ↦ div_pos (half_pos hε) (pow_pos zero_lt_two _) + refine' ⟨f ∘ Encodable.encode, fun i ↦ f0 _, _⟩ rcases hf.summable.comp_injective (@Encodable.encode_injective ι _) with ⟨c, hg⟩ refine' ⟨c, hg, hasSum_le_inj _ (@Encodable.encode_injective ι _) _ _ hg hf⟩ · intro i _ @@ -485,10 +488,10 @@ def posSumOfEncodable {ε : ℝ} (hε : 0 < ε) (ι) [Encodable ι] : #align pos_sum_of_encodable posSumOfEncodable theorem Set.Countable.exists_pos_hasSum_le {ι : Type*} {s : Set ι} (hs : s.Countable) {ε : ℝ} - (hε : 0 < ε) : ∃ ε' : ι → ℝ, (∀ i, 0 < ε' i) ∧ ∃ c, HasSum (fun i : s => ε' i) c ∧ c ≤ ε := by + (hε : 0 < ε) : ∃ ε' : ι → ℝ, (∀ i, 0 < ε' i) ∧ ∃ c, HasSum (fun i : s ↦ ε' i) c ∧ c ≤ ε := by haveI := hs.toEncodable rcases posSumOfEncodable hε s with ⟨f, hf0, ⟨c, hfc, hcε⟩⟩ - refine' ⟨fun i => if h : i ∈ s then f ⟨i, h⟩ else 1, fun i => _, ⟨c, _, hcε⟩⟩ + refine' ⟨fun i ↦ if h : i ∈ s then f ⟨i, h⟩ else 1, fun i ↦ _, ⟨c, _, hcε⟩⟩ · conv_rhs => simp split_ifs exacts [hf0 _, zero_lt_one] @@ -499,10 +502,10 @@ theorem Set.Countable.exists_pos_forall_sum_le {ι : Type*} {s : Set ι} (hs : s (hε : 0 < ε) : ∃ ε' : ι → ℝ, (∀ i, 0 < ε' i) ∧ ∀ t : Finset ι, ↑t ⊆ s → ∑ i in t, ε' i ≤ ε := by rcases hs.exists_pos_hasSum_le hε with ⟨ε', hpos, c, hε'c, hcε⟩ - refine' ⟨ε', hpos, fun t ht => _⟩ + refine' ⟨ε', hpos, fun t ht ↦ _⟩ rw [← sum_subtype_of_mem _ ht] refine' (sum_le_hasSum _ _ hε'c).trans hcε - exact fun _ _ => (hpos _).le + exact fun _ _ ↦ (hpos _).le #align set.countable.exists_pos_forall_sum_le Set.Countable.exists_pos_forall_sum_le namespace NNReal @@ -513,8 +516,8 @@ theorem exists_pos_sum_of_countable {ε : ℝ≥0} (hε : ε ≠ 0) (ι) [Counta obtain ⟨a, a0, aε⟩ := exists_between (pos_iff_ne_zero.2 hε) obtain ⟨ε', hε', c, hc, hcε⟩ := posSumOfEncodable a0 ι exact - ⟨fun i => ⟨ε' i, (hε' i).le⟩, fun i => NNReal.coe_lt_coe.1 <| hε' i, - ⟨c, hasSum_le (fun i => (hε' i).le) hasSum_zero hc⟩, NNReal.hasSum_coe.1 hc, + ⟨fun i ↦ ⟨ε' i, (hε' i).le⟩, fun i ↦ NNReal.coe_lt_coe.1 <| hε' i, + ⟨c, hasSum_le (fun i ↦ (hε' i).le) hasSum_zero hc⟩, NNReal.hasSum_coe.1 hc, aε.trans_le' <| NNReal.coe_le_coe.1 hcε⟩ #align nnreal.exists_pos_sum_of_countable NNReal.exists_pos_sum_of_countable @@ -533,16 +536,16 @@ theorem exists_pos_sum_of_countable {ε : ℝ≥0∞} (hε : ε ≠ 0) (ι) [Cou theorem exists_pos_sum_of_countable' {ε : ℝ≥0∞} (hε : ε ≠ 0) (ι) [Countable ι] : ∃ ε' : ι → ℝ≥0∞, (∀ i, 0 < ε' i) ∧ ∑' i, ε' i < ε := let ⟨δ, δpos, hδ⟩ := exists_pos_sum_of_countable hε ι - ⟨fun i => δ i, fun i => ENNReal.coe_pos.2 (δpos i), hδ⟩ + ⟨fun i ↦ δ i, fun i ↦ ENNReal.coe_pos.2 (δpos i), hδ⟩ #align ennreal.exists_pos_sum_of_countable' ENNReal.exists_pos_sum_of_countable' theorem exists_pos_tsum_mul_lt_of_countable {ε : ℝ≥0∞} (hε : ε ≠ 0) {ι} [Countable ι] (w : ι → ℝ≥0∞) (hw : ∀ i, w i ≠ ∞) : ∃ δ : ι → ℝ≥0, (∀ i, 0 < δ i) ∧ (∑' i, (w i * δ i : ℝ≥0∞)) < ε := by lift w to ι → ℝ≥0 using hw rcases exists_pos_sum_of_countable hε ι with ⟨δ', Hpos, Hsum⟩ - have : ∀ i, 0 < max 1 (w i) := fun i => zero_lt_one.trans_le (le_max_left _ _) - refine' ⟨fun i => δ' i / max 1 (w i), fun i => div_pos (Hpos _) (this i), _⟩ - refine' lt_of_le_of_lt (ENNReal.tsum_le_tsum fun i => _) Hsum + have : ∀ i, 0 < max 1 (w i) := fun i ↦ zero_lt_one.trans_le (le_max_left _ _) + refine' ⟨fun i ↦ δ' i / max 1 (w i), fun i ↦ div_pos (Hpos _) (this i), _⟩ + refine' lt_of_le_of_lt (ENNReal.tsum_le_tsum fun i ↦ _) Hsum rw [coe_div (this i).ne'] refine' mul_le_of_le_div' (mul_le_mul_left' (ENNReal.inv_le_inv.2 _) _) exact coe_le_coe.2 (le_max_right _ _) @@ -556,18 +559,18 @@ end ENNReal theorem factorial_tendsto_atTop : Tendsto Nat.factorial atTop atTop := - tendsto_atTop_atTop_of_monotone Nat.monotone_factorial fun n => ⟨n, n.self_le_factorial⟩ + tendsto_atTop_atTop_of_monotone Nat.monotone_factorial fun n ↦ ⟨n, n.self_le_factorial⟩ #align factorial_tendsto_at_top factorial_tendsto_atTop theorem tendsto_factorial_div_pow_self_atTop : - Tendsto (fun n => n ! / (n : ℝ) ^ n : ℕ → ℝ) atTop (𝓝 0) := + Tendsto (fun n ↦ n ! / (n : ℝ) ^ n : ℕ → ℝ) atTop (𝓝 0) := tendsto_of_tendsto_of_tendsto_of_le_of_le' tendsto_const_nhds (tendsto_const_div_atTop_nhds_0_nat 1) - (eventually_of_forall fun n => + (eventually_of_forall fun n ↦ div_nonneg (mod_cast n.factorial_pos.le) (pow_nonneg (mod_cast n.zero_le) _)) (by - refine' (eventually_gt_atTop 0).mono fun n hn => _ + refine' (eventually_gt_atTop 0).mono fun n hn ↦ _ rcases Nat.exists_eq_succ_of_ne_zero hn.ne.symm with ⟨k, rfl⟩ rw [← prod_range_add_one_eq_factorial, pow_eq_prod_const, div_eq_mul_inv, ← inv_eq_one_div, prod_natCast, Nat.cast_succ, ← prod_inv_distrib, ← prod_mul_distrib, @@ -591,8 +594,8 @@ theorem tendsto_factorial_div_pow_self_atTop : section theorem tendsto_nat_floor_atTop {α : Type*} [LinearOrderedSemiring α] [FloorSemiring α] : - Tendsto (fun x : α => ⌊x⌋₊) atTop atTop := - Nat.floor_mono.tendsto_atTop_atTop fun x => ⟨max 0 (x + 1), by simp [Nat.le_floor_iff]⟩ + Tendsto (fun x : α ↦ ⌊x⌋₊) atTop atTop := + Nat.floor_mono.tendsto_atTop_atTop fun x ↦ ⟨max 0 (x + 1), by simp [Nat.le_floor_iff]⟩ #align tendsto_nat_floor_at_top tendsto_nat_floor_atTop lemma tendsto_nat_floor_mul_atTop {α : Type _} [LinearOrderedSemifield α] [FloorSemiring α] @@ -603,40 +606,40 @@ lemma tendsto_nat_floor_mul_atTop {α : Type _} [LinearOrderedSemifield α] [Flo variable {R : Type*} [TopologicalSpace R] [LinearOrderedField R] [OrderTopology R] [FloorRing R] theorem tendsto_nat_floor_mul_div_atTop {a : R} (ha : 0 ≤ a) : - Tendsto (fun x => (⌊a * x⌋₊ : R) / x) atTop (𝓝 a) := by - have A : Tendsto (fun x : R => a - x⁻¹) atTop (𝓝 (a - 0)) := + Tendsto (fun x ↦ (⌊a * x⌋₊ : R) / x) atTop (𝓝 a) := by + have A : Tendsto (fun x : R ↦ a - x⁻¹) atTop (𝓝 (a - 0)) := tendsto_const_nhds.sub tendsto_inv_atTop_zero rw [sub_zero] at A apply tendsto_of_tendsto_of_tendsto_of_le_of_le' A tendsto_const_nhds - · refine' eventually_atTop.2 ⟨1, fun x hx => _⟩ + · refine' eventually_atTop.2 ⟨1, fun x hx ↦ _⟩ simp only [le_div_iff (zero_lt_one.trans_le hx), _root_.sub_mul, inv_mul_cancel (zero_lt_one.trans_le hx).ne'] have := Nat.lt_floor_add_one (a * x) linarith - · refine' eventually_atTop.2 ⟨1, fun x hx => _⟩ + · refine' eventually_atTop.2 ⟨1, fun x hx ↦ _⟩ rw [div_le_iff (zero_lt_one.trans_le hx)] simp [Nat.floor_le (mul_nonneg ha (zero_le_one.trans hx))] #align tendsto_nat_floor_mul_div_at_top tendsto_nat_floor_mul_div_atTop -theorem tendsto_nat_floor_div_atTop : Tendsto (fun x => (⌊x⌋₊ : R) / x) atTop (𝓝 1) := by +theorem tendsto_nat_floor_div_atTop : Tendsto (fun x ↦ (⌊x⌋₊ : R) / x) atTop (𝓝 1) := by simpa using tendsto_nat_floor_mul_div_atTop (zero_le_one' R) #align tendsto_nat_floor_div_at_top tendsto_nat_floor_div_atTop theorem tendsto_nat_ceil_mul_div_atTop {a : R} (ha : 0 ≤ a) : - Tendsto (fun x => (⌈a * x⌉₊ : R) / x) atTop (𝓝 a) := by - have A : Tendsto (fun x : R => a + x⁻¹) atTop (𝓝 (a + 0)) := + Tendsto (fun x ↦ (⌈a * x⌉₊ : R) / x) atTop (𝓝 a) := by + have A : Tendsto (fun x : R ↦ a + x⁻¹) atTop (𝓝 (a + 0)) := tendsto_const_nhds.add tendsto_inv_atTop_zero rw [add_zero] at A apply tendsto_of_tendsto_of_tendsto_of_le_of_le' tendsto_const_nhds A - · refine' eventually_atTop.2 ⟨1, fun x hx => _⟩ + · refine' eventually_atTop.2 ⟨1, fun x hx ↦ _⟩ rw [le_div_iff (zero_lt_one.trans_le hx)] exact Nat.le_ceil _ - · refine' eventually_atTop.2 ⟨1, fun x hx => _⟩ + · refine' eventually_atTop.2 ⟨1, fun x hx ↦ _⟩ simp [div_le_iff (zero_lt_one.trans_le hx), inv_mul_cancel (zero_lt_one.trans_le hx).ne', (Nat.ceil_lt_add_one (mul_nonneg ha (zero_le_one.trans hx))).le, add_mul] #align tendsto_nat_ceil_mul_div_at_top tendsto_nat_ceil_mul_div_atTop -theorem tendsto_nat_ceil_div_atTop : Tendsto (fun x => (⌈x⌉₊ : R) / x) atTop (𝓝 1) := by +theorem tendsto_nat_ceil_div_atTop : Tendsto (fun x ↦ (⌈x⌉₊ : R) / x) atTop (𝓝 1) := by simpa using tendsto_nat_ceil_mul_div_atTop (zero_le_one' R) #align tendsto_nat_ceil_div_at_top tendsto_nat_ceil_div_atTop diff --git a/Mathlib/Analysis/SpecificLimits/FloorPow.lean b/Mathlib/Analysis/SpecificLimits/FloorPow.lean index 82ea4049d0a68..9be939fc3a6d3 100644 --- a/Mathlib/Analysis/SpecificLimits/FloorPow.lean +++ b/Mathlib/Analysis/SpecificLimits/FloorPow.lean @@ -70,7 +70,7 @@ theorem tendsto_div_of_monotone_of_exists_subseq_tendsto_div (u : ℕ → ℝ) ( let N := Nat.find exN have ncN : n < c N := Nat.find_spec exN have aN : a + 1 ≤ N := by - by_contra' h + by_contra! h have cNM : c N ≤ M := by apply le_max' apply mem_image_of_mem @@ -133,7 +133,7 @@ theorem tendsto_div_of_monotone_of_exists_subseq_tendsto_div (u : ℕ → ℝ) ( let N := Nat.find exN have ncN : n < c N := Nat.find_spec exN have aN : a + 1 ≤ N := by - by_contra' h + by_contra! h have cNM : c N ≤ M := by apply le_max' apply mem_image_of_mem diff --git a/Mathlib/Analysis/SpecificLimits/Normed.lean b/Mathlib/Analysis/SpecificLimits/Normed.lean index 2f9983f5ee720..9a93c7c549190 100644 --- a/Mathlib/Analysis/SpecificLimits/Normed.lean +++ b/Mathlib/Analysis/SpecificLimits/Normed.lean @@ -513,7 +513,7 @@ theorem summable_of_ratio_norm_eventually_le {α : Type*} [SeminormedAddCommGrou · push_neg at hr₀ refine' .of_norm_bounded_eventually_nat 0 summable_zero _ filter_upwards [h] with _ hn - by_contra' h + by_contra! h exact not_lt.mpr (norm_nonneg _) (lt_of_le_of_lt hn <| mul_neg_of_neg_of_pos hr₀ h) #align summable_of_ratio_norm_eventually_le summable_of_ratio_norm_eventually_le diff --git a/Mathlib/Analysis/VonNeumannAlgebra/Basic.lean b/Mathlib/Analysis/VonNeumannAlgebra/Basic.lean index 5d6587b9a329c..966395c0d3d7c 100644 --- a/Mathlib/Analysis/VonNeumannAlgebra/Basic.lean +++ b/Mathlib/Analysis/VonNeumannAlgebra/Basic.lean @@ -143,7 +143,7 @@ theorem coe_commutant (S : VonNeumannAlgebra H) : @[simp] theorem mem_commutant_iff {S : VonNeumannAlgebra H} {z : H →L[ℂ] H} : z ∈ S.commutant ↔ ∀ g ∈ S, g * z = z * g := by - rw [←SetLike.mem_coe, coe_commutant] + rw [← SetLike.mem_coe, coe_commutant] rfl #align von_neumann_algebra.mem_commutant_iff VonNeumannAlgebra.mem_commutant_iff diff --git a/Mathlib/CategoryTheory/Adjunction/Evaluation.lean b/Mathlib/CategoryTheory/Adjunction/Evaluation.lean index c29c77ce3dc11..84fade269826f 100644 --- a/Mathlib/CategoryTheory/Adjunction/Evaluation.lean +++ b/Mathlib/CategoryTheory/Adjunction/Evaluation.lean @@ -64,7 +64,7 @@ def evaluationAdjunctionRight (c : C) : evaluationLeftAdjoint D c ⊣ (evaluatio ext x dsimp ext g - simp only [colimit.ι_desc, Cofan.mk_ι_app, Category.assoc, ←f.naturality, + simp only [colimit.ι_desc, Cofan.mk_ι_app, Category.assoc, ← f.naturality, evaluationLeftAdjoint_obj_map, colimit.ι_desc_assoc, Discrete.functor_obj, Cofan.mk_pt, Discrete.natTrans_app, Category.id_comp] right_inv := fun f => by diff --git a/Mathlib/CategoryTheory/Arrow.lean b/Mathlib/CategoryTheory/Arrow.lean index 1a11c102b42c2..7bddd39f09b92 100644 --- a/Mathlib/CategoryTheory/Arrow.lean +++ b/Mathlib/CategoryTheory/Arrow.lean @@ -231,7 +231,7 @@ instance mono_left [Mono sq] : Mono sq.left where apply CommaMorphism.ext · exact h · rw [Comma.comp_right, Comma.comp_right, this, this, Category.assoc, Category.assoc] - rw [←Arrow.w] + rw [← Arrow.w] simp only [← Category.assoc, h] #align category_theory.arrow.mono_left CategoryTheory.Arrow.mono_left diff --git a/Mathlib/CategoryTheory/Bicategory/Coherence.lean b/Mathlib/CategoryTheory/Bicategory/Coherence.lean index d1e9287f8a210..1ffceaac2c20b 100644 --- a/Mathlib/CategoryTheory/Bicategory/Coherence.lean +++ b/Mathlib/CategoryTheory/Bicategory/Coherence.lean @@ -226,7 +226,7 @@ def normalizeEquiv (a b : B) : Hom a b ≌ Discrete (Path.{v + 1} a b) := injection ih with ih conv => rhs - rw [←ih])) + rw [← ih])) #align category_theory.free_bicategory.normalize_equiv CategoryTheory.FreeBicategory.normalizeEquiv /-- The coherence theorem for bicategories. -/ diff --git a/Mathlib/CategoryTheory/Endofunctor/Algebra.lean b/Mathlib/CategoryTheory/Endofunctor/Algebra.lean index fedceab6f7ef8..28f06066230bc 100644 --- a/Mathlib/CategoryTheory/Endofunctor/Algebra.lean +++ b/Mathlib/CategoryTheory/Endofunctor/Algebra.lean @@ -230,7 +230,8 @@ def strInv : A.1 ⟶ F.obj A.1 := #align category_theory.endofunctor.algebra.initial.str_inv CategoryTheory.Endofunctor.Algebra.Initial.strInv -theorem left_inv' : ⟨strInv h ≫ A.str, by rw [←Category.assoc, F.map_comp, strInv, ←Hom.h]⟩ = 𝟙 A := +theorem left_inv' : + ⟨strInv h ≫ A.str, by rw [← Category.assoc, F.map_comp, strInv, ← Hom.h]⟩ = 𝟙 A := Limits.IsInitial.hom_ext h _ (𝟙 A) #align category_theory.endofunctor.algebra.initial.left_inv' CategoryTheory.Endofunctor.Algebra.Initial.left_inv' @@ -355,7 +356,7 @@ def isoMk (h : V₀.1 ≅ V₁.1) (w : V₀.str ≫ F.map h.hom = h.hom ≫ V₁ inv := { f := h.inv h := by - rw [h.eq_inv_comp, ←Category.assoc, ← w, Category.assoc, ← F.map_comp] + rw [h.eq_inv_comp, ← Category.assoc, ← w, Category.assoc, ← F.map_comp] simp only [Iso.hom_inv_id, Functor.map_id, Category.comp_id] } #align category_theory.endofunctor.coalgebra.iso_mk CategoryTheory.Endofunctor.Coalgebra.isoMk diff --git a/Mathlib/CategoryTheory/Filtered/Basic.lean b/Mathlib/CategoryTheory/Filtered/Basic.lean index bbe37867f6cc9..7e69286333920 100644 --- a/Mathlib/CategoryTheory/Filtered/Basic.lean +++ b/Mathlib/CategoryTheory/Filtered/Basic.lean @@ -454,7 +454,7 @@ theorem bowtie {j₁ j₂ k₁ k₂ : C} (f₁ : j₁ ⟶ k₁) (g₁ : j₁ ⟶ obtain ⟨t, k₁t, k₂t, ht⟩ := span f₁ g₁ obtain ⟨s, ts, hs⟩ := IsFilteredOrEmpty.cocone_maps (f₂ ≫ k₁t) (g₂ ≫ k₂t) simp_rw [Category.assoc] at hs - exact ⟨s, k₁t ≫ ts, k₂t ≫ ts, by simp only [←Category.assoc, ht], hs⟩ + exact ⟨s, k₁t ≫ ts, k₂t ≫ ts, by simp only [← Category.assoc, ht], hs⟩ #align category_theory.is_filtered.bowtie CategoryTheory.IsFiltered.bowtie /-- Given a "tulip" of morphisms @@ -480,7 +480,7 @@ theorem tulip {j₁ j₂ j₃ k₁ k₂ l : C} (f₁ : j₁ ⟶ k₁) (f₂ : j f₁ ≫ α = g₁ ≫ β ∧ f₂ ≫ α = f₃ ≫ γ ∧ f₄ ≫ γ = g₂ ≫ β := by obtain ⟨l', k₁l, k₂l, hl⟩ := span f₂ f₃ obtain ⟨s, ls, l's, hs₁, hs₂⟩ := bowtie g₁ (f₁ ≫ k₁l) g₂ (f₄ ≫ k₂l) - refine' ⟨s, k₁l ≫ l's, ls, k₂l ≫ l's, _, by simp only [←Category.assoc, hl], _⟩ <;> + refine' ⟨s, k₁l ≫ l's, ls, k₂l ≫ l's, _, by simp only [← Category.assoc, hl], _⟩ <;> simp only [hs₁, hs₂, Category.assoc] #align category_theory.is_filtered.tulip CategoryTheory.IsFiltered.tulip diff --git a/Mathlib/CategoryTheory/Functor/Category.lean b/Mathlib/CategoryTheory/Functor/Category.lean index 1c71c587947cf..7b38695892cb1 100644 --- a/Mathlib/CategoryTheory/Functor/Category.lean +++ b/Mathlib/CategoryTheory/Functor/Category.lean @@ -79,11 +79,13 @@ theorem comp_app {F G H : C ⥤ D} (α : F ⟶ G) (β : G ⟶ H) (X : C) : attribute [reassoc] comp_app +@[reassoc] theorem app_naturality {F G : C ⥤ D ⥤ E} (T : F ⟶ G) (X : C) {Y Z : D} (f : Y ⟶ Z) : (F.obj X).map f ≫ (T.app X).app Z = (T.app X).app Y ≫ (G.obj X).map f := (T.app X).naturality f #align category_theory.nat_trans.app_naturality CategoryTheory.NatTrans.app_naturality +@[reassoc] theorem naturality_app {F G : C ⥤ D ⥤ E} (T : F ⟶ G) (Z : D) {X Y : C} (f : X ⟶ Y) : (F.map f).app Z ≫ (T.app Y).app Z = (T.app X).app Z ≫ (G.map f).app Z := congr_fun (congr_arg app (T.naturality f)) Z diff --git a/Mathlib/CategoryTheory/Functor/Currying.lean b/Mathlib/CategoryTheory/Functor/Currying.lean index fe114fdf38424..4ab1dafdfb27f 100644 --- a/Mathlib/CategoryTheory/Functor/Currying.lean +++ b/Mathlib/CategoryTheory/Functor/Currying.lean @@ -54,13 +54,13 @@ def curryObj (F : C × D ⥤ E) : C ⥤ D ⥤ E obj X := { obj := fun Y => F.obj (X, Y) map := fun g => F.map (𝟙 X, g) - map_id := fun Y => by simp only [F.map_id]; rw [←prod_id]; exact F.map_id ⟨X,Y⟩ - map_comp := fun f g => by simp [←F.map_comp]} + map_id := fun Y => by simp only [F.map_id]; rw [← prod_id]; exact F.map_id ⟨X,Y⟩ + map_comp := fun f g => by simp [← F.map_comp]} map f := { app := fun Y => F.map (f, 𝟙 Y) - naturality := fun {Y} {Y'} g => by simp [←F.map_comp] } + naturality := fun {Y} {Y'} g => by simp [← F.map_comp] } map_id := fun X => by ext Y; exact F.map_id _ - map_comp := fun f g => by ext Y; dsimp; simp [←F.map_comp] + map_comp := fun f g => by ext Y; dsimp; simp [← F.map_comp] #align category_theory.curry_obj CategoryTheory.curryObj /-- The currying functor, taking a functor `(C × D) ⥤ E` and producing a functor `C ⥤ (D ⥤ E)`. @@ -88,7 +88,7 @@ def currying : C ⥤ D ⥤ E ≌ C × D ⥤ E := (NatIso.ofComponents fun F => NatIso.ofComponents fun X => NatIso.ofComponents fun Y => Iso.refl _) (NatIso.ofComponents fun F => NatIso.ofComponents (fun X => eqToIso (by simp)) - (by intros X Y f; cases X; cases Y; cases f; dsimp at *; rw [←F.map_comp]; simp)) + (by intros X Y f; cases X; cases Y; cases f; dsimp at *; rw [← F.map_comp]; simp)) #align category_theory.currying CategoryTheory.currying /-- `F.flip` is isomorphic to uncurrying `F`, swapping the variables, and currying. -/ diff --git a/Mathlib/CategoryTheory/Functor/FullyFaithful.lean b/Mathlib/CategoryTheory/Functor/FullyFaithful.lean index 7644e3b3877b2..f675d3967ae16 100644 --- a/Mathlib/CategoryTheory/Functor/FullyFaithful.lean +++ b/Mathlib/CategoryTheory/Functor/FullyFaithful.lean @@ -292,7 +292,7 @@ variable {F F'} def Full.ofIso [Full F] (α : F ≅ F') : Full F' where preimage {X Y} f := F.preimage ((α.app X).hom ≫ f ≫ (α.app Y).inv) - witness f := by simp [←NatIso.naturality_1 α] + witness f := by simp [← NatIso.naturality_1 α] #align category_theory.full.of_iso CategoryTheory.Full.ofIso theorem Faithful.of_iso [Faithful F] (α : F ≅ F') : Faithful F' := diff --git a/Mathlib/CategoryTheory/GradedObject/Trifunctor.lean b/Mathlib/CategoryTheory/GradedObject/Trifunctor.lean new file mode 100644 index 0000000000000..41f6b3ae074c2 --- /dev/null +++ b/Mathlib/CategoryTheory/GradedObject/Trifunctor.lean @@ -0,0 +1,244 @@ +/- +Copyright (c) 2023 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ +import Mathlib.CategoryTheory.GradedObject.Bifunctor +/-! +# The action of trifunctors on graded objects + +Given a trifunctor `F. C₁ ⥤ C₂ ⥤ C₃ ⥤ C₄` and types `I₁`, `I₂` and `I₃`, we define a functor +`GradedObject I₁ C₁ ⥤ GradedObject I₂ C₂ ⥤ GradedObject I₃ C₃ ⥤ GradedObject (I₁ × I₂ × I₃) C₄` +(see `mapTrifunctor`). When we have a map `p : I₁ × I₂ × I₃ → J` and suitable coproducts +exists, we define a functor +`GradedObject I₁ C₁ ⥤ GradedObject I₂ C₂ ⥤ GradedObject I₃ C₃ ⥤ GradedObject J C₄` +(see `mapTrifunctorMap`) which sends graded objects `X₁`, `X₂`, `X₃` to the graded object +which sets `j` to the coproduct of the objects `((F.obj (X₁ i₁)).obj (X₂ i₂)).obj (X₃ i₃)` +for `p ⟨i₁, i₂, i₃⟩ = j`. + +This shall be used in order to construct the associator isomorphism for the monoidal +category structure on `GradedObject I C` induced by a monoidal structure on `C` and +an additive monoid structure on `I` (TODO @joelriou). + +-/ + +namespace CategoryTheory + +open Category Limits + +variable {C₁ C₂ C₃ C₄ : Type*} + [Category C₁] [Category C₂] [Category C₃] [Category C₄] + (F F' : C₁ ⥤ C₂ ⥤ C₃ ⥤ C₄) + +namespace GradedObject + +/-- Auxiliary definition for `mapTrifunctor`. -/ +@[simps] +def mapTrifunctorObj {I₁ : Type*} (X₁ : GradedObject I₁ C₁) (I₂ I₃ : Type*) : + GradedObject I₂ C₂ ⥤ GradedObject I₃ C₃ ⥤ GradedObject (I₁ × I₂ × I₃) C₄ where + obj X₂ := + { obj := fun X₃ x => ((F.obj (X₁ x.1)).obj (X₂ x.2.1)).obj (X₃ x.2.2) + map := fun {X₃ Y₃} φ x => ((F.obj (X₁ x.1)).obj (X₂ x.2.1)).map (φ x.2.2) } + map {X₂ Y₂} φ := + { app := fun X₃ x => ((F.obj (X₁ x.1)).map (φ x.2.1)).app (X₃ x.2.2) } + +/-- Given a trifunctor `F : C₁ ⥤ C₂ ⥤ C₃ ⥤ C₄` and types `I₁`, `I₂`, `I₃`, +this is the obvious functor +`GradedObject I₁ C₁ ⥤ GradedObject I₂ C₂ ⥤ GradedObject I₃ C₃ ⥤ GradedObject (I₁ × I₂ × I₃) C₄`. +-/ +@[simps] +def mapTrifunctor (I₁ I₂ I₃ : Type*) : + GradedObject I₁ C₁ ⥤ GradedObject I₂ C₂ ⥤ GradedObject I₃ C₃ ⥤ + GradedObject (I₁ × I₂ × I₃) C₄ where + obj X₁ := mapTrifunctorObj F X₁ I₂ I₃ + map {X₁ Y₁} φ := + { app := fun X₂ => + { app := fun X₃ x => ((F.map (φ x.1)).app (X₂ x.2.1)).app (X₃ x.2.2) } + naturality := fun {X₂ Y₂} ψ => by + ext X₃ x + dsimp + simp only [← NatTrans.comp_app] + congr 1 + rw [NatTrans.naturality] } + +section + +variable {F F'} + +/-- The natural transformation `mapTrifunctor F I₁ I₂ I₃ ⟶ mapTrifunctor F' I₁ I₂ I₃` +induced by a natural transformation `F ⟶ F` of trifunctors. -/ +@[simps] +def mapTrifunctorMapNatTrans (α : F ⟶ F') (I₁ I₂ I₃ : Type*) : + mapTrifunctor F I₁ I₂ I₃ ⟶ mapTrifunctor F' I₁ I₂ I₃ where + app X₁ := + { app := fun X₂ => + { app := fun X₃ i => ((α.app _).app _).app _ } + naturality := fun {X₂ Y₂} φ => by + ext X₃ ⟨i₁, i₂, i₃⟩ + dsimp + simp only [← NatTrans.comp_app, NatTrans.naturality] } + naturality := fun {X₁ Y₁} φ => by + ext X₂ X₃ ⟨i₁, i₂, i₃⟩ + dsimp + simp only [← NatTrans.comp_app, NatTrans.naturality] + +/-- The natural isomorphism `mapTrifunctor F I₁ I₂ I₃ ≅ mapTrifunctor F' I₁ I₂ I₃` +induced by a natural isomorphism `F ≅ F` of trifunctors. -/ +@[simps] +def mapTrifunctorMapIso (e : F ≅ F') (I₁ I₂ I₃ : Type*) : + mapTrifunctor F I₁ I₂ I₃ ≅ mapTrifunctor F' I₁ I₂ I₃ where + hom := mapTrifunctorMapNatTrans e.hom I₁ I₂ I₃ + inv := mapTrifunctorMapNatTrans e.inv I₁ I₂ I₃ + hom_inv_id := by + ext X₁ X₂ X₃ ⟨i₁, i₂, i₃⟩ + dsimp + simp only [← NatTrans.comp_app, e.hom_inv_id, NatTrans.id_app] + inv_hom_id := by + ext X₁ X₂ X₃ ⟨i₁, i₂, i₃⟩ + dsimp + simp only [← NatTrans.comp_app, e.inv_hom_id, NatTrans.id_app] + +end + +section + +variable {I₁ I₂ I₃ J : Type*} (p : I₁ × I₂ × I₃ → J) + +/-- Given a trifunctor `F : C₁ ⥤ C₂ ⥤ C₃ ⥤ C₃`, graded objects `X₁ : GradedObject I₁ C₁`, +`X₂ : GradedObject I₂ C₂`, `X₃ : GradedObject I₃ C₃`, and a map `p : I₁ × I₂ × I₃ → J`, +this is the `J`-graded object sending `j` to the coproduct of +`((F.obj (X₁ i₁)).obj (X₂ i₂)).obj (X₃ i₃)` for `p ⟨i₁, i₂, i₃⟩ = k`. -/ +noncomputable def mapTrifunctorMapObj (X₁ : GradedObject I₁ C₁) (X₂ : GradedObject I₂ C₂) + (X₃ : GradedObject I₃ C₃) + [HasMap ((((mapTrifunctor F I₁ I₂ I₃).obj X₁).obj X₂).obj X₃) p] : + GradedObject J C₄ := + ((((mapTrifunctor F I₁ I₂ I₃).obj X₁).obj X₂).obj X₃).mapObj p + +/-- The obvious inclusion +`((F.obj (X₁ i₁)).obj (X₂ i₂)).obj (X₃ i₃) ⟶ mapTrifunctorMapObj F p X₁ X₂ X₃ j` when +`p ⟨i₁, i₂, i₃⟩ = j`. -/ +noncomputable def ιMapTrifunctorMapObj (X₁ : GradedObject I₁ C₁) (X₂ : GradedObject I₂ C₂) + (X₃ : GradedObject I₃ C₃) (i₁ : I₁) (i₂ : I₂) (i₃ : I₃) (j : J) (h : p ⟨i₁, i₂, i₃⟩ = j) + [HasMap ((((mapTrifunctor F I₁ I₂ I₃).obj X₁).obj X₂).obj X₃) p] : + ((F.obj (X₁ i₁)).obj (X₂ i₂)).obj (X₃ i₃) ⟶ mapTrifunctorMapObj F p X₁ X₂ X₃ j := + ((((mapTrifunctor F I₁ I₂ I₃).obj X₁).obj X₂).obj X₃).ιMapObj p ⟨i₁, i₂, i₃⟩ j h + +/-- The maps `mapTrifunctorMapObj F p X₁ X₂ X₃ ⟶ mapTrifunctorMapObj F p Y₁ Y₂ Y₃` which +express the functoriality of `mapTrifunctorMapObj`, see `mapTrifunctorMap` -/ +noncomputable def mapTrifunctorMapMap {X₁ Y₁ : GradedObject I₁ C₁} (f₁ : X₁ ⟶ Y₁) + {X₂ Y₂ : GradedObject I₂ C₂} (f₂ : X₂ ⟶ Y₂) + {X₃ Y₃ : GradedObject I₃ C₃} (f₃ : X₃ ⟶ Y₃) + [HasMap ((((mapTrifunctor F I₁ I₂ I₃).obj X₁).obj X₂).obj X₃) p] + [HasMap ((((mapTrifunctor F I₁ I₂ I₃).obj Y₁).obj Y₂).obj Y₃) p] : + mapTrifunctorMapObj F p X₁ X₂ X₃ ⟶ mapTrifunctorMapObj F p Y₁ Y₂ Y₃ := + GradedObject.mapMap ((((mapTrifunctor F I₁ I₂ I₃).map f₁).app X₂).app X₃ ≫ + (((mapTrifunctor F I₁ I₂ I₃).obj Y₁).map f₂).app X₃ ≫ + (((mapTrifunctor F I₁ I₂ I₃).obj Y₁).obj Y₂).map f₃) p + +@[reassoc (attr := simp)] +lemma ι_mapTrifunctorMapMap {X₁ Y₁ : GradedObject I₁ C₁} (f₁ : X₁ ⟶ Y₁) + {X₂ Y₂ : GradedObject I₂ C₂} (f₂ : X₂ ⟶ Y₂) + {X₃ Y₃ : GradedObject I₃ C₃} (f₃ : X₃ ⟶ Y₃) + [HasMap ((((mapTrifunctor F I₁ I₂ I₃).obj X₁).obj X₂).obj X₃) p] + [HasMap ((((mapTrifunctor F I₁ I₂ I₃).obj Y₁).obj Y₂).obj Y₃) p] + (i₁ : I₁) (i₂ : I₂) (i₃ : I₃) (j : J) (h : p ⟨i₁, i₂, i₃⟩ = j) : + ιMapTrifunctorMapObj F p X₁ X₂ X₃ i₁ i₂ i₃ j h ≫ mapTrifunctorMapMap F p f₁ f₂ f₃ j = + ((F.map (f₁ i₁)).app (X₂ i₂)).app (X₃ i₃) ≫ + ((F.obj (Y₁ i₁)).map (f₂ i₂)).app (X₃ i₃) ≫ + ((F.obj (Y₁ i₁)).obj (Y₂ i₂)).map (f₃ i₃) ≫ + ιMapTrifunctorMapObj F p Y₁ Y₂ Y₃ i₁ i₂ i₃ j h := by + dsimp only [ιMapTrifunctorMapObj, mapTrifunctorMapMap] + rw [ι_mapMap] + dsimp + rw [assoc, assoc] + +@[ext] +lemma mapTrifunctorMapObj_ext {X₁ : GradedObject I₁ C₁} {X₂ : GradedObject I₂ C₂} + {X₃ : GradedObject I₃ C₃} {Y : C₄} (j : J) + [HasMap ((((mapTrifunctor F I₁ I₂ I₃).obj X₁).obj X₂).obj X₃) p] + {φ φ' : mapTrifunctorMapObj F p X₁ X₂ X₃ j ⟶ Y} + (h : ∀ (i₁ : I₁) (i₂ : I₂) (i₃ : I₃) (h : p ⟨i₁, i₂, i₃⟩ = j), + ιMapTrifunctorMapObj F p X₁ X₂ X₃ i₁ i₂ i₃ j h ≫ φ = + ιMapTrifunctorMapObj F p X₁ X₂ X₃ i₁ i₂ i₃ j h ≫ φ') : φ = φ' := by + apply mapObj_ext + rintro ⟨i₁, i₂, i₃⟩ hi + apply h + +instance (X₁ : GradedObject I₁ C₁) (X₂ : GradedObject I₂ C₂) (X₃ : GradedObject I₃ C₃) + [h : HasMap ((((mapTrifunctor F I₁ I₂ I₃).obj X₁).obj X₂).obj X₃) p] : + HasMap (((mapTrifunctorObj F X₁ I₂ I₃).obj X₂).obj X₃) p := h + +/-- Given a trifunctor `F : C₁ ⥤ C₂ ⥤ C₃ ⥤ C₄`, a map `p : I₁ × I₂ × I₃ → J`, and +graded objects `X₁ : GradedObject I₁ C₁`, `X₂ : GradedObject I₂ C₂` and `X₃ : GradedObject I₃ C₃`, +this is the `J`-graded object sending `j` to the coproduct of +`((F.obj (X₁ i₁)).obj (X₂ i₂)).obj (X₃ i₃)` for `p ⟨i₁, i₂, i₃⟩ = j`. -/ +@[simps] +noncomputable def mapTrifunctorMapFunctorObj (X₁ : GradedObject I₁ C₁) + [∀ X₂ X₃, HasMap ((((mapTrifunctor F I₁ I₂ I₃).obj X₁).obj X₂).obj X₃) p] : + GradedObject I₂ C₂ ⥤ GradedObject I₃ C₃ ⥤ GradedObject J C₄ where + obj X₂ := + { obj := fun X₃ => mapTrifunctorMapObj F p X₁ X₂ X₃ + map := fun {X₃ Y₃} φ => mapTrifunctorMapMap F p (𝟙 X₁) (𝟙 X₂) φ + map_id := fun X₃ => by + ext j i₁ i₂ i₃ h + dsimp + simp only [ι_mapTrifunctorMapMap, categoryOfGradedObjects_id, Functor.map_id, + NatTrans.id_app, id_comp, comp_id] + map_comp := fun {X₃ Y₃ Z₃} φ ψ => by + ext j i₁ i₂ i₃ h + dsimp + simp only [ι_mapTrifunctorMapMap, categoryOfGradedObjects_id, Functor.map_id, + NatTrans.id_app, categoryOfGradedObjects_comp, Functor.map_comp, assoc, id_comp, + ι_mapTrifunctorMapMap_assoc] } + map {X₂ Y₂} φ := + { app := fun X₃ => mapTrifunctorMapMap F p (𝟙 X₁) φ (𝟙 X₃) + naturality := fun {X₃ Y₃} ψ => by + ext j i₁ i₂ i₃ h + dsimp + simp only [ι_mapTrifunctorMapMap_assoc, categoryOfGradedObjects_id, Functor.map_id, + NatTrans.id_app, ι_mapTrifunctorMapMap, id_comp, NatTrans.naturality_assoc] } + map_id X₂ := by + ext X₃ j i₁ i₂ i₃ h + dsimp + simp only [ι_mapTrifunctorMapMap, categoryOfGradedObjects_id, Functor.map_id, + NatTrans.id_app, id_comp, comp_id] + map_comp {X₂ Y₂ Z₂} φ ψ := by + ext X₃ j i₁ i₂ i₃ + dsimp + simp only [ι_mapTrifunctorMapMap, categoryOfGradedObjects_id, Functor.map_id, + NatTrans.id_app, categoryOfGradedObjects_comp, Functor.map_comp, NatTrans.comp_app, + id_comp, assoc, ι_mapTrifunctorMapMap_assoc] + +/-- Given a trifunctor `F : C₁ ⥤ C₂ ⥤ C₃ ⥤ C₄` and a map `p : I₁ × I₂ × I₃ → J`, +this is the functor +`GradedObject I₁ C₁ ⥤ GradedObject I₂ C₂ ⥤ GradedObject I₃ C₃ ⥤ GradedObject J C₄` +sending `X₁ : GradedObject I₁ C₁`, `X₂ : GradedObject I₂ C₂` and `X₃ : GradedObject I₃ C₃` +to the `J`-graded object sending `j` to the coproduct of +`((F.obj (X₁ i₁)).obj (X₂ i₂)).obj (X₃ i₃)` for `p ⟨i₁, i₂, i₃⟩ = j`. -/ +@[simps] +noncomputable def mapTrifunctorMap + [∀ X₁ X₂ X₃, HasMap ((((mapTrifunctor F I₁ I₂ I₃).obj X₁).obj X₂).obj X₃) p] : + GradedObject I₁ C₁ ⥤ GradedObject I₂ C₂ ⥤ GradedObject I₃ C₃ ⥤ GradedObject J C₄ where + obj X₁ := mapTrifunctorMapFunctorObj F p X₁ + map := fun {X₁ Y₁} φ => + { app := fun X₂ => + { app := fun X₃ => mapTrifunctorMapMap F p φ (𝟙 X₂) (𝟙 X₃) + naturality := fun {X₃ Y₃} φ => by + dsimp + ext j i₁ i₂ i₃ h + dsimp + simp only [ι_mapTrifunctorMapMap_assoc, categoryOfGradedObjects_id, Functor.map_id, + NatTrans.id_app, ι_mapTrifunctorMapMap, id_comp, NatTrans.naturality_assoc] } + naturality := fun {X₂ Y₂} ψ => by + ext X₃ j + dsimp + ext i₁ i₂ i₃ h + simp only [ι_mapTrifunctorMapMap_assoc, categoryOfGradedObjects_id, Functor.map_id, + NatTrans.id_app, ι_mapTrifunctorMapMap, id_comp, + NatTrans.naturality_app_assoc] } + +end + +end GradedObject + +end CategoryTheory diff --git a/Mathlib/CategoryTheory/Groupoid/FreeGroupoid.lean b/Mathlib/CategoryTheory/Groupoid/FreeGroupoid.lean index 124b4692f9510..214a7d4fe997f 100644 --- a/Mathlib/CategoryTheory/Groupoid/FreeGroupoid.lean +++ b/Mathlib/CategoryTheory/Groupoid/FreeGroupoid.lean @@ -112,7 +112,7 @@ theorem congr_comp_reverse {X Y : Paths <| Quiver.Symmetrify V} (p : X ⟶ Y) : -- category notation change Quotient.CompClosure redStep (q ≫ Quiver.Path.reverse q) (Quiver.Path.cons q f ≫ (Quiver.Hom.toPath (Quiver.reverse f)) ≫ (Quiver.Path.reverse q)) - simp only [←Category.assoc] at this ⊢ + simp only [← Category.assoc] at this ⊢ exact this · exact ih #align category_theory.groupoid.free.congr_comp_reverse CategoryTheory.Groupoid.Free.congr_comp_reverse diff --git a/Mathlib/CategoryTheory/Limits/IsLimit.lean b/Mathlib/CategoryTheory/Limits/IsLimit.lean index a7363494ac057..89efa5e44924e 100644 --- a/Mathlib/CategoryTheory/Limits/IsLimit.lean +++ b/Mathlib/CategoryTheory/Limits/IsLimit.lean @@ -377,7 +377,7 @@ def conePointsIsoOfEquivalence {F : J ⥤ C} {s : Cone F} {G : K ⥤ C} {t : Con dsimp simp only [Limits.Cone.whisker_π, Limits.Cones.postcompose_obj_π, fac, whiskerLeft_app, assoc, id_comp, invFunIdAssoc_hom_app, fac_assoc, NatTrans.comp_app] - rw [counit_app_functor, ←Functor.comp_map] + rw [counit_app_functor, ← Functor.comp_map] have l : NatTrans.app w.hom j = NatTrans.app w.hom (Prefunctor.obj (𝟭 J).toPrefunctor j) := by dsimp rw [l,w.hom.naturality] diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Pullbacks.lean b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Pullbacks.lean index 2b1a793b37a1a..3392ea32c2732 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Pullbacks.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Pullbacks.lean @@ -61,7 +61,7 @@ def isLimitMapConePullbackConeEquiv : /-- The property of preserving pullbacks expressed in terms of binary fans. -/ def isLimitPullbackConeMapOfIsLimit [PreservesLimit (cospan f g) G] (l : IsLimit (PullbackCone.mk h k comm)) : - have : G.map h ≫ G.map f = G.map k ≫ G.map g := by rw [←G.map_comp,←G.map_comp,comm] + have : G.map h ≫ G.map f = G.map k ≫ G.map g := by rw [← G.map_comp, ← G.map_comp,comm] IsLimit (PullbackCone.mk (G.map h) (G.map k) this) := isLimitMapConePullbackConeEquiv G comm (PreservesLimit.preserves l) #align category_theory.limits.is_limit_pullback_cone_map_of_is_limit CategoryTheory.Limits.isLimitPullbackConeMapOfIsLimit @@ -69,7 +69,7 @@ def isLimitPullbackConeMapOfIsLimit [PreservesLimit (cospan f g) G] /-- The property of reflecting pullbacks expressed in terms of binary fans. -/ def isLimitOfIsLimitPullbackConeMap [ReflectsLimit (cospan f g) G] (l : IsLimit (PullbackCone.mk (G.map h) (G.map k) (show G.map h ≫ G.map f = G.map k ≫ G.map g - from by simp only [←G.map_comp,comm]))) : IsLimit (PullbackCone.mk h k comm) := + from by simp only [← G.map_comp,comm]))) : IsLimit (PullbackCone.mk h k comm) := ReflectsLimit.reflects ((isLimitMapConePullbackConeEquiv G comm).symm l) #align category_theory.limits.is_limit_of_is_limit_pullback_cone_map CategoryTheory.Limits.isLimitOfIsLimitPullbackConeMap @@ -79,7 +79,7 @@ variable (f g) [PreservesLimit (cospan f g) G] morphisms of the pullback cone is a limit. -/ def isLimitOfHasPullbackOfPreservesLimit [i : HasPullback f g] : have : G.map pullback.fst ≫ G.map f = G.map pullback.snd ≫ G.map g := by - simp only [←G.map_comp, pullback.condition]; + simp only [← G.map_comp, pullback.condition]; IsLimit (PullbackCone.mk (G.map (@pullback.fst _ _ _ _ _ f g i)) (G.map pullback.snd) this) := isLimitPullbackConeMapOfIsLimit G _ (pullbackIsPullback f g) #align category_theory.limits.is_limit_of_has_pullback_of_preserves_limit CategoryTheory.Limits.isLimitOfHasPullbackOfPreservesLimit @@ -168,14 +168,14 @@ def isColimitMapCoconePushoutCoconeEquiv : def isColimitPushoutCoconeMapOfIsColimit [PreservesColimit (span f g) G] (l : IsColimit (PushoutCocone.mk h k comm)) : IsColimit (PushoutCocone.mk (G.map h) (G.map k) (show G.map f ≫ G.map h = G.map g ≫ G.map k - from by simp only [←G.map_comp,comm] )) := + from by simp only [← G.map_comp,comm] )) := isColimitMapCoconePushoutCoconeEquiv G comm (PreservesColimit.preserves l) #align category_theory.limits.is_colimit_pushout_cocone_map_of_is_colimit CategoryTheory.Limits.isColimitPushoutCoconeMapOfIsColimit /-- The property of reflecting pushouts expressed in terms of binary cofans. -/ def isColimitOfIsColimitPushoutCoconeMap [ReflectsColimit (span f g) G] (l : IsColimit (PushoutCocone.mk (G.map h) (G.map k) (show G.map f ≫ G.map h = - G.map g ≫ G.map k from by simp only [←G.map_comp,comm]))) : + G.map g ≫ G.map k from by simp only [← G.map_comp,comm]))) : IsColimit (PushoutCocone.mk h k comm) := ReflectsColimit.reflects ((isColimitMapCoconePushoutCoconeEquiv G comm).symm l) #align category_theory.limits.is_colimit_of_is_colimit_pushout_cocone_map CategoryTheory.Limits.isColimitOfIsColimitPushoutCoconeMap diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Equalizers.lean b/Mathlib/CategoryTheory/Limits/Shapes/Equalizers.lean index d5f7fe6545f16..5d6e69590ac39 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Equalizers.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Equalizers.lean @@ -1228,7 +1228,7 @@ variable {C f g} /-- The fork obtained by postcomposing an equalizer fork with a monomorphism is an equalizer. -/ def isEqualizerCompMono {c : Fork f g} (i : IsLimit c) {Z : C} (h : Y ⟶ Z) [hm : Mono h] : have : Fork.ι c ≫ f ≫ h = Fork.ι c ≫ g ≫ h := by - simp only [←Category.assoc] + simp only [← Category.assoc] exact congrArg (· ≫ h) c.condition; IsLimit (Fork.ofι c.ι (by simp [this]) : Fork (f ≫ h) (g ≫ h)) := Fork.IsLimit.mk' _ fun s => diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Images.lean b/Mathlib/CategoryTheory/Limits/Shapes/Images.lean index 4550733585534..a6f08ee7a8fa3 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Images.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Images.lean @@ -595,7 +595,7 @@ instance hasImage_iso_comp [IsIso f] [HasImage g] : HasImage (f ≫ g) := lift_fac := fun F' => by dsimp have : (MonoFactorisation.ofIsoComp f F').m = F'.m := rfl - rw [←this,image.lift_fac (MonoFactorisation.ofIsoComp f F')] } } + rw [← this,image.lift_fac (MonoFactorisation.ofIsoComp f F')] } } #align category_theory.limits.has_image_iso_comp CategoryTheory.Limits.hasImage_iso_comp /-- `image.preComp f g` is an isomorphism when `f` is an isomorphism @@ -622,7 +622,7 @@ instance hasImage_comp_iso [HasImage f] [IsIso g] : HasImage (f ≫ g) := { lift := fun F' => image.lift F'.ofCompIso lift_fac := fun F' => by rw [← Category.comp_id (image.lift (MonoFactorisation.ofCompIso F') ≫ F'.m), - ←IsIso.inv_hom_id g,← Category.assoc] + ← IsIso.inv_hom_id g,← Category.assoc] refine congrArg (· ≫ g) ?_ have : (image.lift (MonoFactorisation.ofCompIso F') ≫ F'.m) ≫ inv g = image.lift (MonoFactorisation.ofCompIso F') ≫ diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Kernels.lean b/Mathlib/CategoryTheory/Limits/Shapes/Kernels.lean index 66547d9db4ae3..c6b7dbbe0a91c 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Kernels.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Kernels.lean @@ -533,7 +533,7 @@ def IsKernel.isoKernel {Z : C} (l : Z ⟶ X) {s : KernelFork f} (hs : IsLimit s) Cones.ext i.symm fun j => by cases j · exact (Iso.eq_inv_comp i).2 h - · dsimp; rw[←h]; simp + · dsimp; rw[← h]; simp #align category_theory.limits.is_kernel.iso_kernel CategoryTheory.Limits.IsKernel.isoKernel /-- If `i` is an isomorphism such that `i.hom ≫ kernel.ι f = l`, then `l` is a kernel of `f`. -/ @@ -798,7 +798,7 @@ abbrev cokernel.map {X' Y' : C} (f' : X' ⟶ Y') [HasCokernel f'] (p : X ⟶ X') (w : f ≫ q = p ≫ f') : cokernel f ⟶ cokernel f' := cokernel.desc f (q ≫ cokernel.π f') (by have : f ≫ q ≫ π f' = p ≫ f' ≫ π f' := by - simp only [←Category.assoc] + simp only [← Category.assoc] apply congrArg (· ≫ π f') w simp [this]) #align category_theory.limits.cokernel.map CategoryTheory.Limits.cokernel.map @@ -1114,7 +1114,7 @@ def IsCokernel.cokernelIso {Z : C} (l : Y ⟶ Z) {s : CokernelCofork f} (hs : Is IsColimit.ofIsoColimit hs <| Cocones.ext i fun j => by cases j - · dsimp; rw [←h]; simp + · dsimp; rw [← h]; simp · exact h #align category_theory.limits.is_cokernel.cokernel_iso CategoryTheory.Limits.IsCokernel.cokernelIso diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Pullbacks.lean b/Mathlib/CategoryTheory/Limits/Shapes/Pullbacks.lean index 18bbc9caa8806..341f3d913d230 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Pullbacks.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Pullbacks.lean @@ -635,7 +635,7 @@ theorem mono_snd_of_is_pullback_of_mono {t : PullbackCone f g} (ht : IsLimit t) theorem mono_fst_of_is_pullback_of_mono {t : PullbackCone f g} (ht : IsLimit t) [Mono g] : Mono t.fst := by refine ⟨fun {W} h k i => IsLimit.hom_ext ht i ?_⟩ - rw [← cancel_mono g, Category.assoc, Category.assoc, ←condition] + rw [← cancel_mono g, Category.assoc, Category.assoc, ← condition] have := congrArg (· ≫ f) i; dsimp at this rwa [Category.assoc, Category.assoc] at this #align category_theory.limits.pullback_cone.mono_fst_of_is_pullback_of_mono CategoryTheory.Limits.PullbackCone.mono_fst_of_is_pullback_of_mono diff --git a/Mathlib/CategoryTheory/Limits/Shapes/StrongEpi.lean b/Mathlib/CategoryTheory/Limits/Shapes/StrongEpi.lean index 0c84ff1a5f841..51eed021dd96f 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/StrongEpi.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/StrongEpi.lean @@ -132,7 +132,7 @@ theorem strongMono_of_strongMono [StrongMono (f ≫ g)] : StrongMono f := constructor intro u v sq have h₀ : u ≫ f ≫ g = z ≫ v ≫ g := by - rw [←Category.assoc, eq_whisker sq.w, Category.assoc] + rw [← Category.assoc, eq_whisker sq.w, Category.assoc] exact CommSq.HasLift.mk' ⟨(CommSq.mk h₀).lift, by simp, by simp [← cancel_epi z, sq.w]⟩ } #align category_theory.strong_mono_of_strong_mono CategoryTheory.strongMono_of_strongMono diff --git a/Mathlib/CategoryTheory/Limits/Shapes/ZeroMorphisms.lean b/Mathlib/CategoryTheory/Limits/Shapes/ZeroMorphisms.lean index bf3472b2f108f..857332e2f7a4f 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/ZeroMorphisms.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/ZeroMorphisms.lean @@ -111,7 +111,7 @@ theorem ext (I J : HasZeroMorphisms C) : I = J := by apply I.zero_comp X (J.zero Y Y).zero have that : (I.zero X Y).zero ≫ (J.zero Y Y).zero = (J.zero X Y).zero := by apply J.comp_zero (I.zero X Y).zero Y - rw[←this,←that] + rw[← this, ← that] #align category_theory.limits.has_zero_morphisms.ext CategoryTheory.Limits.HasZeroMorphisms.ext instance : Subsingleton (HasZeroMorphisms C) := @@ -467,7 +467,7 @@ def isoOfIsIsomorphicZero {X : C} (P : IsIsomorphic X 0) : X ≅ 0 where inv := 0 hom_inv_id := by cases' P with P - rw [←P.hom_inv_id,←Category.id_comp P.inv] + rw [← P.hom_inv_id, ← Category.id_comp P.inv] apply Eq.symm simp only [id_comp, Iso.hom_inv_id, comp_zero] apply (idZeroEquivIsoZero X).invFun P diff --git a/Mathlib/CategoryTheory/Monoidal/Category.lean b/Mathlib/CategoryTheory/Monoidal/Category.lean index b488953de56ab..8a5e479c7a4e8 100644 --- a/Mathlib/CategoryTheory/Monoidal/Category.lean +++ b/Mathlib/CategoryTheory/Monoidal/Category.lean @@ -482,7 +482,7 @@ abbrev ofTensorHom [MonoidalCategoryStruct C] tensorHom (rightUnitor X).hom (𝟙 Y) := by aesop_cat) : MonoidalCategory C where - tensorHom_def := by intros; simp [← id_tensorHom, ←tensorHom_id, ← tensor_comp] + tensorHom_def := by intros; simp [← id_tensorHom, ← tensorHom_id, ← tensor_comp] whiskerLeft_id := by intros; simp [← id_tensorHom, ← tensor_id] id_whiskerRight := by intros; simp [← tensorHom_id, tensor_id] pentagon := pentagon diff --git a/Mathlib/CategoryTheory/Monoidal/Preadditive.lean b/Mathlib/CategoryTheory/Monoidal/Preadditive.lean index a5f122a1aa6f8..2cd3cb0892697 100644 --- a/Mathlib/CategoryTheory/Monoidal/Preadditive.lean +++ b/Mathlib/CategoryTheory/Monoidal/Preadditive.lean @@ -97,7 +97,7 @@ theorem tensor_sum {P Q R S : C} {J : Type*} (s : Finset J) (f : P ⟶ Q) (g : J rw [← tensor_id_comp_id_tensor] let tQ := (((tensoringLeft C).obj Q).mapAddHom : (R ⟶ S) →+ _) change _ ≫ tQ _ = _ - rw [tQ.map_sum, Preadditive.comp_sum] + rw [map_sum, Preadditive.comp_sum] dsimp [Functor.mapAddHom] simp only [tensor_id_comp_id_tensor] #align category_theory.tensor_sum CategoryTheory.tensor_sum @@ -107,7 +107,7 @@ theorem sum_tensor {P Q R S : C} {J : Type*} (s : Finset J) (f : P ⟶ Q) (g : J rw [← tensor_id_comp_id_tensor] let tQ := (((tensoringRight C).obj P).mapAddHom : (R ⟶ S) →+ _) change tQ _ ≫ _ = _ - rw [tQ.map_sum, Preadditive.sum_comp] + rw [map_sum, Preadditive.sum_comp] dsimp [Functor.mapAddHom] simp only [tensor_id_comp_id_tensor] #align category_theory.sum_tensor CategoryTheory.sum_tensor diff --git a/Mathlib/CategoryTheory/Monoidal/Transport.lean b/Mathlib/CategoryTheory/Monoidal/Transport.lean index 5c674a6ffa602..19a28a34f189d 100644 --- a/Mathlib/CategoryTheory/Monoidal/Transport.lean +++ b/Mathlib/CategoryTheory/Monoidal/Transport.lean @@ -117,25 +117,25 @@ abbrev induced [MonoidalCategoryStruct D] (F : D ⥤ C) [Faithful F] simp only [Functor.map_comp, fData.tensorHom_eq, Functor.map_id, fData.leftUnitor_eq, Iso.trans_assoc, Iso.trans_hom, Iso.symm_hom, tensorIso_hom, Iso.refl_hom, assoc, Iso.hom_inv_id_assoc, id_tensor_comp_tensor_id_assoc, Iso.cancel_iso_inv_left] - rw [←this, ←assoc, ←tensor_comp, id_comp, comp_id] + rw [← this, ← assoc, ← tensor_comp, id_comp, comp_id] rightUnitor_naturality {X Y : D} f := F.map_injective <| by have := rightUnitor_naturality (F.map f) simp only [Functor.map_comp, fData.tensorHom_eq, Functor.map_id, fData.rightUnitor_eq, Iso.trans_assoc, Iso.trans_hom, Iso.symm_hom, tensorIso_hom, Iso.refl_hom, assoc, Iso.hom_inv_id_assoc, tensor_id_comp_id_tensor_assoc, Iso.cancel_iso_inv_left] - rw [←this, ←assoc, ←tensor_comp, id_comp, comp_id] + rw [← this, ← assoc, ← tensor_comp, id_comp, comp_id] associator_naturality {X₁ X₂ X₃ Y₁ Y₂ Y₃} f₁ f₂ f₃ := F.map_injective <| by have := associator_naturality (F.map f₁) (F.map f₂) (F.map f₃) simp [fData.associator_eq, fData.tensorHom_eq] - simp_rw [←assoc, ←tensor_comp, assoc, Iso.hom_inv_id, ←assoc] + simp_rw [← assoc, ← tensor_comp, assoc, Iso.hom_inv_id, ← assoc] congr 1 - conv_rhs => rw [←comp_id (F.map f₁), ←id_comp (F.map f₁)] + conv_rhs => rw [← comp_id (F.map f₁), ← id_comp (F.map f₁)] simp only [tensor_comp] simp only [tensor_id, comp_id, assoc, tensor_hom_inv_id_assoc, id_comp] - slice_rhs 2 3 => rw [←this] + slice_rhs 2 3 => rw [← this] simp only [← assoc, Iso.inv_hom_id, comp_id] congr 2 - simp_rw [←tensor_comp, id_comp] + simp_rw [← tensor_comp, id_comp] /-- diff --git a/Mathlib/CategoryTheory/Over.lean b/Mathlib/CategoryTheory/Over.lean index 815dfe2934201..c7054c1c3c4f7 100644 --- a/Mathlib/CategoryTheory/Over.lean +++ b/Mathlib/CategoryTheory/Over.lean @@ -243,7 +243,7 @@ The converse of `CategoryTheory.Over.mono_of_mono_left`. instance mono_left_of_mono {f g : Over X} (k : f ⟶ g) [Mono k] : Mono k.left := by refine' ⟨fun { Y : T } l m a => _⟩ let l' : mk (m ≫ f.hom) ⟶ f := homMk l (by - dsimp; rw [← Over.w k, ←Category.assoc, congrArg (· ≫ g.hom) a, Category.assoc]) + dsimp; rw [← Over.w k, ← Category.assoc, congrArg (· ≫ g.hom) a, Category.assoc]) suffices l' = (homMk m : mk (m ≫ f.hom) ⟶ f) by apply congrArg CommaMorphism.left this rw [← cancel_mono k] ext diff --git a/Mathlib/CategoryTheory/Preadditive/AdditiveFunctor.lean b/Mathlib/CategoryTheory/Preadditive/AdditiveFunctor.lean index 684fdcade21fd..30313a7ab95fd 100644 --- a/Mathlib/CategoryTheory/Preadditive/AdditiveFunctor.lean +++ b/Mathlib/CategoryTheory/Preadditive/AdditiveFunctor.lean @@ -98,9 +98,9 @@ theorem map_zsmul {X Y : C} {f : X ⟶ Y} {r : ℤ} : F.map (r • f) = r • F. open BigOperators @[simp] -theorem map_sum {X Y : C} {α : Type*} (f : α → (X ⟶ Y)) (s : Finset α) : +nonrec theorem map_sum {X Y : C} {α : Type*} (f : α → (X ⟶ Y)) (s : Finset α) : F.map (∑ a in s, f a) = ∑ a in s, F.map (f a) := - (F.mapAddHom : (X ⟶ Y) →+ _).map_sum f s + map_sum F.mapAddHom f s #align category_theory.functor.map_sum CategoryTheory.Functor.map_sum variable {F} diff --git a/Mathlib/CategoryTheory/Preadditive/Biproducts.lean b/Mathlib/CategoryTheory/Preadditive/Biproducts.lean index cbde774362322..ead97e98263d2 100644 --- a/Mathlib/CategoryTheory/Preadditive/Biproducts.lean +++ b/Mathlib/CategoryTheory/Preadditive/Biproducts.lean @@ -811,7 +811,7 @@ def Biprod.isoElim (f : X₁ ⊞ X₂ ≅ Y₁ ⊞ Y₂) [IsIso (biprod.inl ≫ theorem Biprod.column_nonzero_of_iso {W X Y Z : C} (f : W ⊞ X ⟶ Y ⊞ Z) [IsIso f] : 𝟙 W = 0 ∨ biprod.inl ≫ f ≫ biprod.fst ≠ 0 ∨ biprod.inl ≫ f ≫ biprod.snd ≠ 0 := by - by_contra' h + by_contra! h rcases h with ⟨nz, a₁, a₂⟩ set x := biprod.inl ≫ f ≫ inv f ≫ biprod.fst have h₁ : x = 𝟙 W := by simp diff --git a/Mathlib/CategoryTheory/Preadditive/Opposite.lean b/Mathlib/CategoryTheory/Preadditive/Opposite.lean index 31198585b8fac..60bfb95405a3b 100644 --- a/Mathlib/CategoryTheory/Preadditive/Opposite.lean +++ b/Mathlib/CategoryTheory/Preadditive/Opposite.lean @@ -73,7 +73,7 @@ def unopHom (X Y : Cᵒᵖ) : (X ⟶ Y) →+ (Opposite.unop Y ⟶ Opposite.unop @[simp] theorem unop_sum (X Y : Cᵒᵖ) {ι : Type*} (s : Finset ι) (f : ι → (X ⟶ Y)) : (s.sum f).unop = s.sum fun i => (f i).unop := - (unopHom X Y).map_sum _ _ + map_sum (unopHom X Y) _ _ #align category_theory.unop_sum CategoryTheory.unop_sum /-- `op` induces morphisms of monoids on hom groups of a preadditive category -/ @@ -85,7 +85,7 @@ def opHom (X Y : C) : (X ⟶ Y) →+ (Opposite.op Y ⟶ Opposite.op X) := @[simp] theorem op_sum (X Y : C) {ι : Type*} (s : Finset ι) (f : ι → (X ⟶ Y)) : (s.sum f).op = s.sum fun i => (f i).op := - (opHom X Y).map_sum _ _ + map_sum (opHom X Y) _ _ #align category_theory.op_sum CategoryTheory.op_sum variable {D : Type*} [Category D] [Preadditive D] diff --git a/Mathlib/CategoryTheory/Shift/Induced.lean b/Mathlib/CategoryTheory/Shift/Induced.lean index 500790854268b..7fb619c2cff24 100644 --- a/Mathlib/CategoryTheory/Shift/Induced.lean +++ b/Mathlib/CategoryTheory/Shift/Induced.lean @@ -156,7 +156,7 @@ noncomputable def induced : HasShift D A := slice_lhs 6 7 => erw [← Functor.map_comp, ← Functor.map_comp, Iso.inv_hom_id_app, (s m₂).map_id, (s m₃).map_id] - erw [Category.comp_id, ←NatTrans.naturality_assoc, reassoc_of% eq, + erw [Category.comp_id, ← NatTrans.naturality_assoc, reassoc_of% eq, dcongr_arg (fun a => (i a).hom.app X) (add_assoc m₁ m₂ m₃).symm] simp only [Functor.comp_obj, eqToHom_map, eqToHom_app, NatTrans.naturality_assoc, Induced.add_hom_app_obj, Functor.comp_map, Category.assoc, Iso.inv_hom_id_app_assoc, @@ -235,7 +235,7 @@ def Functor.CommShift.ofInduced : simp only [isoAdd_hom_app, Iso.symm_hom, shiftFunctorAdd_inv_app_obj_of_induced, shiftFunctor_of_induced] erw [← Functor.map_comp_assoc, Iso.inv_hom_id_app, Functor.map_id, - Category.id_comp, Iso.inv_hom_id_app_assoc, ←F.map_comp_assoc, Iso.hom_inv_id_app, + Category.id_comp, Iso.inv_hom_id_app_assoc, ← F.map_comp_assoc, Iso.hom_inv_id_app, F.map_id, Category.id_comp] } lemma Functor.commShiftIso_eq_ofInduced (a : A) : diff --git a/Mathlib/CategoryTheory/Sites/Coherent.lean b/Mathlib/CategoryTheory/Sites/Coherent.lean index e56c144e87bf1..af9d0562dcee9 100644 --- a/Mathlib/CategoryTheory/Sites/Coherent.lean +++ b/Mathlib/CategoryTheory/Sites/Coherent.lean @@ -116,7 +116,7 @@ theorem isSheaf_yoneda_obj (W : C) : Presieve.IsSheaf (coherentTopology C) (yone rw [isSheaf_coherent] intro X α _ Y π H have h_colim := isColimitOfEffectiveEpiFamilyStruct Y π H.effectiveEpiFamily.some - rw [←Sieve.generateFamily_eq] at h_colim + rw [← Sieve.generateFamily_eq] at h_colim intro x hx let x_ext := Presieve.FamilyOfElements.sieveExtend x have hx_ext := Presieve.FamilyOfElements.Compatible.sieveExtend hx diff --git a/Mathlib/CategoryTheory/Sites/Coverage.lean b/Mathlib/CategoryTheory/Sites/Coverage.lean index 034dadcd169d8..d6c8f96b5c5e4 100644 --- a/Mathlib/CategoryTheory/Sites/Coverage.lean +++ b/Mathlib/CategoryTheory/Sites/Coverage.lean @@ -103,7 +103,7 @@ lemma isSheafFor_of_factorsThru (h : ∀ ⦃Y : C⦄ ⦃f : Y ⟶ X⦄, T f → ∃ (R : Presieve Y), R.IsSeparatedFor P ∧ R.FactorsThruAlong S f): T.IsSheafFor P := by - simp only [←Presieve.isSeparatedFor_and_exists_isAmalgamation_iff_isSheafFor] at * + simp only [← Presieve.isSeparatedFor_and_exists_isAmalgamation_iff_isSheafFor] at * choose W i e h1 h2 using H refine ⟨?_, fun x hx => ?_⟩ · intro x y₁ y₂ h₁ h₂ diff --git a/Mathlib/CategoryTheory/Sites/Sheafification.lean b/Mathlib/CategoryTheory/Sites/Sheafification.lean index 342a24eb8b7d9..d610e91cea37a 100644 --- a/Mathlib/CategoryTheory/Sites/Sheafification.lean +++ b/Mathlib/CategoryTheory/Sites/Sheafification.lean @@ -196,7 +196,7 @@ theorem toPlus_apply {X : C} {P : Cᵒᵖ ⥤ D} (S : J.Cover X) (x : Meq P S) ( dsimp only [toPlus, plusObj] delta Cover.toMultiequalizer dsimp [mk] - erw [←comp_apply] + erw [← comp_apply] rw [ι_colimMap_assoc, colimit.ι_pre, comp_apply, comp_apply] dsimp only [Functor.op] let e : (J.pullback I.f).obj (unop (op S)) ⟶ ⊤ := homOfLE (OrderTop.le_top _) @@ -206,7 +206,7 @@ theorem toPlus_apply {X : C} {P : Cᵒᵖ ⥤ D} (S : J.Cover X) (x : Meq P S) ( apply Concrete.multiequalizer_ext intro i dsimp [diagram] - rw [←comp_apply, ←comp_apply, ←comp_apply, Multiequalizer.lift_ι, Multiequalizer.lift_ι, + rw [← comp_apply, ← comp_apply, ← comp_apply, Multiequalizer.lift_ι, Multiequalizer.lift_ι, Multiequalizer.lift_ι] erw [Meq.equiv_symm_eq_apply] let RR : S.Relation := @@ -224,7 +224,7 @@ theorem toPlus_eq_mk {X : C} {P : Cᵒᵖ ⥤ D} (x : P.obj (op X)) : apply congr_arg apply (Meq.equiv P ⊤).injective ext i - rw [Meq.equiv_apply, Equiv.apply_symm_apply, ←comp_apply, Multiequalizer.lift_ι] + rw [Meq.equiv_apply, Equiv.apply_symm_apply, ← comp_apply, Multiequalizer.lift_ι] rfl #align category_theory.grothendieck_topology.plus.to_plus_eq_mk CategoryTheory.GrothendieckTopology.Plus.toPlus_eq_mk diff --git a/Mathlib/CategoryTheory/Sites/Surjective.lean b/Mathlib/CategoryTheory/Sites/Surjective.lean index 1a299dbfadeb1..053e385d95f3b 100644 --- a/Mathlib/CategoryTheory/Sites/Surjective.lean +++ b/Mathlib/CategoryTheory/Sites/Surjective.lean @@ -113,7 +113,7 @@ theorem isLocallySurjective_of_iso {F G : Cᵒᵖ ⥤ A} (f : F ⟶ G) [IsIso f] apply isLocallySurjective_of_surjective intro U apply Function.Bijective.surjective - rw [← isIso_iff_bijective, ←forget_map_eq_coe] + rw [← isIso_iff_bijective, ← forget_map_eq_coe] infer_instance #align category_theory.is_locally_surjective_of_iso CategoryTheory.isLocallySurjective_of_iso diff --git a/Mathlib/CategoryTheory/StructuredArrow.lean b/Mathlib/CategoryTheory/StructuredArrow.lean index 43a1bac5d7e4a..92e16f96a0a10 100644 --- a/Mathlib/CategoryTheory/StructuredArrow.lean +++ b/Mathlib/CategoryTheory/StructuredArrow.lean @@ -298,7 +298,7 @@ noncomputable def isEquivalencePre (S : D) (F : B ⥤ C) (G : C ⥤ D) [IsEquiva def post (S : C) (F : B ⥤ C) (G : C ⥤ D) : StructuredArrow S F ⥤ StructuredArrow (G.obj S) (F ⋙ G) where obj X := StructuredArrow.mk (G.map X.hom) - map f := StructuredArrow.homMk f.right (by simp [Functor.comp_map, ←G.map_comp, ← f.w]) + map f := StructuredArrow.homMk f.right (by simp [Functor.comp_map, ← G.map_comp, ← f.w]) #align category_theory.structured_arrow.post CategoryTheory.StructuredArrow.post instance (S : C) (F : B ⥤ C) (G : C ⥤ D) : Faithful (post S F G) where @@ -626,7 +626,7 @@ noncomputable def isEquivalencePre (F : B ⥤ C) (G : C ⥤ D) (S : D) [IsEquiva def post (F : B ⥤ C) (G : C ⥤ D) (S : C) : CostructuredArrow F S ⥤ CostructuredArrow (F ⋙ G) (G.obj S) where obj X := CostructuredArrow.mk (G.map X.hom) - map f := CostructuredArrow.homMk f.left (by simp [Functor.comp_map, ←G.map_comp, ← f.w]) + map f := CostructuredArrow.homMk f.left (by simp [Functor.comp_map, ← G.map_comp, ← f.w]) #align category_theory.costructured_arrow.post CategoryTheory.CostructuredArrow.post instance (F : B ⥤ C) (G : C ⥤ D) (S : C) : Faithful (post F G S) where diff --git a/Mathlib/CategoryTheory/Yoneda.lean b/Mathlib/CategoryTheory/Yoneda.lean index 091c67c66427d..0546dfc31cd1b 100644 --- a/Mathlib/CategoryTheory/Yoneda.lean +++ b/Mathlib/CategoryTheory/Yoneda.lean @@ -351,7 +351,7 @@ def yonedaLemma : yonedaPairing C ≅ yonedaEvaluation C where simp only [yonedaEvaluation] ext dsimp - erw [Category.id_comp, ←FunctorToTypes.naturality] + erw [Category.id_comp, ← FunctorToTypes.naturality] simp only [Category.comp_id, yoneda_obj_map] } inv := { app := fun F x => @@ -366,7 +366,7 @@ def yonedaLemma : yonedaPairing C ≅ yonedaEvaluation C where simp only [yoneda] ext dsimp - rw [←FunctorToTypes.naturality X.snd Y.snd f.snd, FunctorToTypes.map_comp_apply] } + rw [← FunctorToTypes.naturality X.snd Y.snd f.snd, FunctorToTypes.map_comp_apply] } hom_inv_id := by ext dsimp diff --git a/Mathlib/Combinatorics/Catalan.lean b/Mathlib/Combinatorics/Catalan.lean index 7585d7cda142e..1bd5673e2c062 100644 --- a/Mathlib/Combinatorics/Catalan.lean +++ b/Mathlib/Combinatorics/Catalan.lean @@ -98,7 +98,8 @@ private theorem gosper_trick {n i : ℕ} (h : i ≤ n) : mod_cast Nat.succ_mul_centralBinom_succ (n - i) simp only [gosperCatalan] push_cast - rw [show n + 1 - i = n - i + 1 by rw [Nat.add_comm (n - i) 1, ←(Nat.add_sub_assoc h 1), add_comm]] + rw [show n + 1 - i = n - i + 1 by rw [Nat.add_comm (n - i) 1, ← (Nat.add_sub_assoc h 1), + add_comm]] rw [h₁, h₂, h₃, h₄] field_simp ring diff --git a/Mathlib/Combinatorics/Colex.lean b/Mathlib/Combinatorics/Colex.lean index 0edc2182ed773..6fb4153f4a31a 100644 --- a/Mathlib/Combinatorics/Colex.lean +++ b/Mathlib/Combinatorics/Colex.lean @@ -119,7 +119,7 @@ private lemma trans_aux (hst : toColex s ≤ toColex t) (htu : toColex t ≤ toC private lemma antisymm_aux (hst : toColex s ≤ toColex t) (hts : toColex t ≤ toColex s) : s ⊆ t := by intro a has - by_contra' hat + by_contra! hat have ⟨_b, hb₁, hb₂, _⟩ := trans_aux hst hts has hat exact hb₂ hb₁ @@ -232,8 +232,8 @@ lemma le_iff_sdiff_subset_lowerClosure {s t : Colex α} : /-- The colexigraphic order is insensitive to removing the same elements from both sets. -/ lemma toColex_sdiff_le_toColex_sdiff (hus : u ⊆ s) (hut : u ⊆ t) : toColex (s \ u) ≤ toColex (t \ u) ↔ toColex s ≤ toColex t := by - simp_rw [toColex_le_toColex, ←and_imp, ←and_assoc, ←mem_sdiff, sdiff_sdiff_sdiff_cancel_right hus, - sdiff_sdiff_sdiff_cancel_right hut] + simp_rw [toColex_le_toColex, ← and_imp, ← and_assoc, ← mem_sdiff, + sdiff_sdiff_sdiff_cancel_right hus, sdiff_sdiff_sdiff_cancel_right hut] /-- The colexigraphic order is insensitive to removing the same elements from both sets. -/ lemma toColex_sdiff_lt_toColex_sdiff (hus : u ⊆ s) (hut : u ⊆ t) : @@ -271,7 +271,7 @@ private lemma max_mem_aux {s t : Colex α} (hst : s ≠ t) : (ofColex s ∆ ofCo lemma toColex_lt_toColex_iff_exists_forall_lt : toColex s < toColex t ↔ ∃ a ∈ t, a ∉ s ∧ ∀ b ∈ s, b ∉ t → b < a := by - rw [←not_le, toColex_le_toColex, not_forall] + rw [← not_le, toColex_le_toColex, not_forall] simp only [not_forall, not_exists, not_and, not_le, exists_prop, exists_and_left] lemma lt_iff_exists_forall_lt {s t : Colex α} : @@ -333,8 +333,8 @@ def IsInitSeg (𝒜 : Finset (Finset α)) (r : ℕ) : Prop := -/ lemma IsInitSeg.total (h₁ : IsInitSeg 𝒜₁ r) (h₂ : IsInitSeg 𝒜₂ r) : 𝒜₁ ⊆ 𝒜₂ ∨ 𝒜₂ ⊆ 𝒜₁ := by classical - simp_rw [←sdiff_eq_empty_iff_subset, ←not_nonempty_iff_eq_empty] - by_contra' h + simp_rw [← sdiff_eq_empty_iff_subset, ← not_nonempty_iff_eq_empty] + by_contra! h have ⟨⟨s, hs⟩, t, ht⟩ := h rw [mem_sdiff] at hs ht obtain hst | hst | hts := trichotomous_of (α := Colex α) (· < ·) (toColex s) (toColex t) @@ -401,7 +401,7 @@ lemma geomSum_ofColex_strictMono (hn : 2 ≤ n) : StrictMono fun s ↦ ∑ k in rintro ⟨s⟩ ⟨t⟩ hst rw [toColex_lt_toColex_iff_exists_forall_lt] at hst obtain ⟨a, hat, has, ha⟩ := hst - rw [←sum_sdiff_lt_sum_sdiff] + rw [← sum_sdiff_lt_sum_sdiff] exact (Nat.geomSum_lt hn $ by simpa).trans_le <| single_le_sum (fun _ _ ↦ by positivity) <| mem_sdiff.2 ⟨hat, has⟩ diff --git a/Mathlib/Combinatorics/Composition.lean b/Mathlib/Combinatorics/Composition.lean index abbd4cd22bdcc..644d7802b21b2 100644 --- a/Mathlib/Combinatorics/Composition.lean +++ b/Mathlib/Combinatorics/Composition.lean @@ -332,7 +332,7 @@ theorem sizeUpTo_index_le (j : Fin n) : c.sizeUpTo (c.index j) ≤ j := by set i := c.index j push_neg at H have i_pos : (0 : ℕ) < i := by - by_contra' i_pos + by_contra! i_pos revert H simp [nonpos_iff_eq_zero.1 i_pos, c.sizeUpTo_zero] let i₁ := (i : ℕ).pred @@ -835,7 +835,7 @@ def compositionAsSetEquiv (n : ℕ) : CompositionAsSet n ≃ Finset (Fin (n - 1) · rw [add_comm] at this contradiction · cases' h with w h; cases' h with h₁ h₂ - rw [←Fin.ext_iff] at h₂ + rw [← Fin.ext_iff] at h₂ rwa [h₂] · intro h apply Or.inr diff --git a/Mathlib/Combinatorics/Partition.lean b/Mathlib/Combinatorics/Partition.lean index 4c5731e8b737a..be8c2d9e85ea9 100644 --- a/Mathlib/Combinatorics/Partition.lean +++ b/Mathlib/Combinatorics/Partition.lean @@ -94,7 +94,7 @@ def ofSums (n : ℕ) (l : Multiset ℕ) (hl : l.sum = n) : Partition n where parts_pos hi := (of_mem_filter hi).bot_lt parts_sum := by have lz : (l.filter (· = 0)).sum = 0 := by simp [sum_eq_zero_iff] - rwa [←filter_add_not (· = 0) l, sum_add, lz, zero_add] at hl + rwa [← filter_add_not (· = 0) l, sum_add, lz, zero_add] at hl #align nat.partition.of_sums Nat.Partition.ofSums /-- A `Multiset ℕ` induces a partition on its sum. -/ diff --git a/Mathlib/Combinatorics/Quiver/Symmetric.lean b/Mathlib/Combinatorics/Quiver/Symmetric.lean index bd1503d8b3b32..93484db807053 100644 --- a/Mathlib/Combinatorics/Quiver/Symmetric.lean +++ b/Mathlib/Combinatorics/Quiver/Symmetric.lean @@ -75,7 +75,7 @@ theorem reverse_inj [h : HasInvolutiveReverse V] {a b : V} theorem eq_reverse_iff [h : HasInvolutiveReverse V] {a b : V} (f : a ⟶ b) (g : b ⟶ a) : f = reverse g ↔ reverse f = g := by - rw [←reverse_inj, reverse_reverse] + rw [← reverse_inj, reverse_reverse] #align quiver.eq_reverse_iff Quiver.eq_reverse_iff section MapReverse diff --git a/Mathlib/Combinatorics/SetFamily/Compression/Down.lean b/Mathlib/Combinatorics/SetFamily/Compression/Down.lean index 75d8d5c2e9885..b61299d7e64d6 100644 --- a/Mathlib/Combinatorics/SetFamily/Compression/Down.lean +++ b/Mathlib/Combinatorics/SetFamily/Compression/Down.lean @@ -142,7 +142,7 @@ lemma memberSubfamily_image_insert (h𝒜 : ∀ s ∈ 𝒜, a ∉ s) : simp only [mem_memberSubfamily, mem_image] refine ⟨?_, fun hs ↦ ⟨⟨s, hs, rfl⟩, h𝒜 _ hs⟩⟩ rintro ⟨⟨t, ht, hts⟩, hs⟩ - rwa [←insert_erase_invOn.2.injOn (h𝒜 _ ht) hs hts] + rwa [← insert_erase_invOn.2.injOn (h𝒜 _ ht) hs hts] @[simp] lemma nonMemberSubfamily_image_insert : (𝒜.image <| insert a).nonMemberSubfamily a = ∅ := by simp [eq_empty_iff_forall_not_mem] @@ -182,7 +182,7 @@ lemma memberFamily_induction_on {p : Finset (Finset α) → Prop} clear_value u induction' u using Finset.induction with a u _ ih generalizing 𝒜 · simp_rw [subset_empty] at hu - rw [←subset_singleton_iff', subset_singleton_iff] at hu + rw [← subset_singleton_iff', subset_singleton_iff] at hu obtain rfl | rfl := hu <;> assumption refine subfamily a (ih _ ?_) (ih _ ?_) · simp only [mem_nonMemberSubfamily, and_imp] @@ -211,7 +211,7 @@ protected lemma family_induction_on {p : Finset (Finset α) → Prop} (subfamily : ∀ (a : α) ⦃𝒜 : Finset (Finset α)⦄, p (𝒜.filter (a ∉ ·)) → p (𝒜.filter (a ∈ ·)) → p 𝒜) : p 𝒜 := by refine memberFamily_induction_on 𝒜 empty singleton_empty fun a 𝒜 h𝒜₀ h𝒜₁ ↦ subfamily a h𝒜₀ ?_ - rw [←image_insert_memberSubfamily] + rw [← image_insert_memberSubfamily] exact image_insert _ (by simp) h𝒜₁ end Finset diff --git a/Mathlib/Combinatorics/SetFamily/Compression/UV.lean b/Mathlib/Combinatorics/SetFamily/Compression/UV.lean index 9313c3b949c97..f9370d6f7256d 100644 --- a/Mathlib/Combinatorics/SetFamily/Compression/UV.lean +++ b/Mathlib/Combinatorics/SetFamily/Compression/UV.lean @@ -205,7 +205,7 @@ theorem compress_mem_compression_of_mem_compression (ha : a ∈ 𝓒 u v s) : theorem compression_idem (u v : α) (s : Finset α) : 𝓒 u v (𝓒 u v s) = 𝓒 u v s := by have h : filter (compress u v · ∉ 𝓒 u v s) (𝓒 u v s) = ∅ := filter_false_of_mem fun a ha h ↦ h <| compress_mem_compression_of_mem_compression ha - rw [compression, image_filter, h, image_empty, ←h] + rw [compression, image_filter, h, image_empty, ← h] exact filter_union_filter_neg_eq _ (compression u v s) #align uv.compression_idem UV.compression_idem @@ -213,7 +213,7 @@ theorem compression_idem (u v : α) (s : Finset α) : 𝓒 u v (𝓒 u v s) = @[simp] theorem card_compression (u v : α) (s : Finset α) : (𝓒 u v s).card = s.card := by rw [compression, card_disjoint_union compress_disjoint, image_filter, - card_image_of_injOn compress_injOn, ←card_disjoint_union (disjoint_filter_filter_neg s _ _), + card_image_of_injOn compress_injOn, ← card_disjoint_union (disjoint_filter_filter_neg s _ _), filter_union_filter_neg_eq] #align uv.card_compression UV.card_compression diff --git a/Mathlib/Combinatorics/SetFamily/FourFunctions.lean b/Mathlib/Combinatorics/SetFamily/FourFunctions.lean index df56819e27edf..f848fba0256fb 100644 --- a/Mathlib/Combinatorics/SetFamily/FourFunctions.lean +++ b/Mathlib/Combinatorics/SetFamily/FourFunctions.lean @@ -116,7 +116,7 @@ lemma le_collapse_of_mem (ha : a ∉ s) (hf : 0 ≤ f) (hts : t = s) (ht : t ∈ lemma le_collapse_of_insert_mem (ha : a ∉ s) (hf : 0 ≤ f) (hts : t = insert a s) (ht : t ∈ 𝒜) : f t ≤ collapse 𝒜 a f s := by - rw [collapse_eq ha, ←hts, if_pos ht] + rw [collapse_eq ha, ← hts, if_pos ht] split_ifs · exact le_add_of_nonneg_left $ hf _ · rw [zero_add] @@ -152,14 +152,14 @@ lemma collapse_modular (hu : a ∉ u) (h₁ : 0 ≤ f₁) (h₂ : 0 ≤ f₂) (h · rw [add_zero, add_mul] refine' (add_le_add (h ‹_› ‹_›) $ h ‹_› ‹_›).trans _ rw [collapse_of_mem ‹_› (union_mem_sups ‹_› ‹_›) (union_mem_sups ‹_› ‹_›) rfl - (insert_union _ _ _), insert_inter_of_not_mem ‹_›, ←mul_add] + (insert_union _ _ _), insert_inter_of_not_mem ‹_›, ← mul_add] exact mul_le_mul_of_nonneg_right (le_collapse_of_mem ‹_› h₃ rfl $ inter_mem_infs ‹_› ‹_›) $ add_nonneg (h₄ _) $ h₄ _ · rw [zero_add, add_mul] refine' (add_le_add (h ‹_› ‹_›) $ h ‹_› ‹_›).trans _ rw [collapse_of_mem ‹_› (inter_mem_infs ‹_› ‹_›) (inter_mem_infs ‹_› ‹_›) (inter_insert_of_not_mem ‹_›) (insert_inter_distrib _ _ _).symm, union_insert, - insert_union_distrib, ←add_mul] + insert_union_distrib, ← add_mul] exact mul_le_mul_of_nonneg_left (le_collapse_of_insert_mem ‹_› h₄ (insert_union_distrib _ _ _).symm $ union_mem_sups ‹_› ‹_›) $ add_nonneg (h₃ _) $ h₃ _ · rw [add_zero, mul_zero] @@ -168,7 +168,7 @@ lemma collapse_modular (hu : a ∉ u) (h₁ : 0 ≤ f₁) (h₂ : 0 ≤ f₂) (h split_ifs · refine' (add_le_add (h ‹_› ‹_›) $ h ‹_› ‹_›).trans _ rw [collapse_of_mem ‹_› (union_mem_sups ‹_› ‹_›) (union_mem_sups ‹_› ‹_›) rfl - (union_insert _ _ _), inter_insert_of_not_mem ‹_›, ←mul_add] + (union_insert _ _ _), inter_insert_of_not_mem ‹_›, ← mul_add] exact mul_le_mul_of_nonneg_right (le_collapse_of_mem ‹_› h₃ rfl $ inter_mem_infs ‹_› ‹_›) $ add_nonneg (h₄ _) $ h₄ _ · rw [mul_zero, add_zero] @@ -187,8 +187,8 @@ lemma collapse_modular (hu : a ∉ u) (h₁ : 0 ≤ f₁) (h₂ : 0 ≤ f₂) (h · refine' (add_le_add (h ‹_› ‹_›) $ h ‹_› ‹_›).trans _ rw [collapse_of_mem ‹_› (inter_mem_infs ‹_› ‹_›) (inter_mem_infs ‹_› ‹_›) (insert_inter_of_not_mem ‹_›) (insert_inter_distrib _ _ _).symm, - insert_inter_of_not_mem ‹_›, ←insert_inter_distrib, insert_union, insert_union_distrib, - ←add_mul] + insert_inter_of_not_mem ‹_›, ← insert_inter_distrib, insert_union, insert_union_distrib, + ← add_mul] exact mul_le_mul_of_nonneg_left (le_collapse_of_insert_mem ‹_› h₄ (insert_union_distrib _ _ _).symm $ union_mem_sups ‹_› ‹_›) $ add_nonneg (h₃ _) $ h₃ _ · rw [mul_zero, add_zero] @@ -210,7 +210,7 @@ lemma sum_collapse (h𝒜 : 𝒜 ⊆ (insert a u).powerset) (hu : a ∉ u) : _ = ∑ s in u.powerset ∩ 𝒜, f s + ∑ s in u.powerset.image (insert a) ∩ 𝒜, f s := ?_ _ = ∑ s in u.powerset ∩ 𝒜, f s + ∑ s in ((insert a u).powerset \ u.powerset) ∩ 𝒜, f s := ?_ _ = ∑ s in 𝒜, f s := ?_ - · rw [←sum_ite_mem, ←sum_ite_mem, sum_image, ←sum_add_distrib] + · rw [← sum_ite_mem, ← sum_ite_mem, sum_image, ← sum_add_distrib] · exact sum_congr rfl fun s hs ↦ collapse_eq (not_mem_mono (mem_powerset.1 hs) hu) _ _ · exact (insert_erase_invOn.2.injOn).mono fun s hs ↦ not_mem_mono (mem_powerset.1 hs) hu · congr with s @@ -221,7 +221,7 @@ lemma sum_collapse (h𝒜 : 𝒜 ⊆ (insert a u).powerset) (hu : a ∉ u) : · rw [insert_erase (erase_ne_self.1 fun hs ↦ ?_)] rw [hs] at h exact h.2 h.1 - · rw [←sum_union (disjoint_sdiff_self_right.mono inf_le_left inf_le_left), ←inter_distrib_right, + · rw [← sum_union (disjoint_sdiff_self_right.mono inf_le_left inf_le_left), ← inter_distrib_right, union_sdiff_of_subset (powerset_mono.2 $ subset_insert _ _), inter_eq_right.2 h𝒜] /-- The **Four Functions Theorem** on a powerset algebra. See `four_functions_theorem` for the @@ -278,13 +278,13 @@ lemma four_functions_theorem (h₁ : 0 ≤ f₁) (h₂ : 0 ≤ f₂) (h₃ : 0 (extend g (f₃ ∘ (↑)) 0) (extend g (f₄ ∘ (↑)) 0) (extend_nonneg (fun _ ↦ h₁ _) le_rfl) (extend_nonneg (fun _ ↦ h₂ _) le_rfl) (extend_nonneg (fun _ ↦ h₃ _) le_rfl) (extend_nonneg (fun _ ↦ h₄ _) le_rfl) ?_ (s'.map ⟨g, hg⟩) (t'.map ⟨g, hg⟩) - simpa only [←hs', ←ht', ←map_sups, ←map_infs, sum_map, Embedding.coeFn_mk, hg.extend_apply] + simpa only [← hs', ← ht', ← map_sups, ← map_infs, sum_map, Embedding.coeFn_mk, hg.extend_apply] using this rintro s t classical obtain ⟨a, rfl⟩ | hs := em (∃ a, g a = s) · obtain ⟨b, rfl⟩ | ht := em (∃ b, g b = t) - · simp_rw [←sup_eq_union, ←inf_eq_inter, ←map_sup, ←map_inf, hg.extend_apply] + · simp_rw [← sup_eq_union, ← inf_eq_inter, ← map_sup, ← map_inf, hg.extend_apply] exact h _ _ · simpa [extend_apply' _ _ _ ht] using mul_nonneg (extend_nonneg (fun a : L ↦ h₃ a) le_rfl _) (extend_nonneg (fun a : L ↦ h₄ a) le_rfl _) @@ -330,7 +330,7 @@ lemma fkg (hμ₀ : 0 ≤ μ) (hf₀ : 0 ≤ f) (hg₀ : 0 ≤ g) (hf : Monotone refine' four_functions_theorem_univ (μ * f) (μ * g) μ _ (mul_nonneg hμ₀ hf₀) (mul_nonneg hμ₀ hg₀) hμ₀ (mul_nonneg hμ₀ $ mul_nonneg hf₀ hg₀) (fun a b ↦ _) dsimp - rw [mul_mul_mul_comm, ←mul_assoc (μ (a ⊓ b))] + rw [mul_mul_mul_comm, ← mul_assoc (μ (a ⊓ b))] exact mul_le_mul (hμ _ _) (mul_le_mul (hf le_sup_left) (hg le_sup_right) (hg₀ _) $ hf₀ _) (mul_nonneg (hf₀ _) $ hg₀ _) $ mul_nonneg (hμ₀ _) $ hμ₀ _ @@ -348,10 +348,10 @@ lemma Finset.le_card_diffs_mul_card_diffs (s t : Finset α) : · rintro s t simp_rw [map_eq_image] exact image_image₂_distrib fun a b ↦ rfl - simpa [←card_compls (_ ⊻ _), ←map_sup, ←map_inf, ←this] using + simpa [← card_compls (_ ⊻ _), ← map_sup, ← map_inf, ← this] using (s.map ⟨_, liftLatticeHom_injective⟩).le_card_infs_mul_card_sups (t.map ⟨_, liftLatticeHom_injective⟩)ᶜˢ /-- The **Marica-Schönheim Inequality**. -/ lemma Finset.card_le_card_diffs (s : Finset α) : s.card ≤ (s \\ s).card := - le_of_pow_le_pow 2 (zero_le _) two_pos $ by simpa [←sq] using s.le_card_diffs_mul_card_diffs s + le_of_pow_le_pow 2 (zero_le _) two_pos $ by simpa [← sq] using s.le_card_diffs_mul_card_diffs s diff --git a/Mathlib/Combinatorics/SetFamily/Intersecting.lean b/Mathlib/Combinatorics/SetFamily/Intersecting.lean index b966a23ef1c40..7a854e03d9e32 100644 --- a/Mathlib/Combinatorics/SetFamily/Intersecting.lean +++ b/Mathlib/Combinatorics/SetFamily/Intersecting.lean @@ -183,7 +183,7 @@ theorem Intersecting.is_max_iff_card_eq (hs : (s : Set α).Intersecting) : image_eq_preimage_of_inverse compl_compl compl_compl] refine' eq_univ_of_forall fun a => _ simp_rw [mem_union, mem_preimage] - by_contra' ha + by_contra! ha refine' s.ne_insert_of_not_mem _ ha.1 (h _ _ <| s.subset_insert _) rw [coe_insert] refine' hs.insert _ fun b hb hab => ha.2 <| (hs.isUpperSet' h) hab.le_compl_left hb diff --git a/Mathlib/Combinatorics/SetFamily/Shadow.lean b/Mathlib/Combinatorics/SetFamily/Shadow.lean index f87fb428d4fe9..65d27fdd28aa5 100644 --- a/Mathlib/Combinatorics/SetFamily/Shadow.lean +++ b/Mathlib/Combinatorics/SetFamily/Shadow.lean @@ -100,12 +100,12 @@ theorem erase_mem_shadow (hs : s ∈ 𝒜) (ha : a ∈ s) : erase s a ∈ ∂ See also `Finset.mem_shadow_iff_exists_mem_card_add_one`. -/ lemma mem_shadow_iff_exists_sdiff : t ∈ ∂ 𝒜 ↔ ∃ s ∈ 𝒜, t ⊆ s ∧ (s \ t).card = 1 := by - simp_rw [mem_shadow_iff, ←covby_iff_card_sdiff_eq_one, covby_iff_exists_erase, eq_comm] + simp_rw [mem_shadow_iff, ← covby_iff_card_sdiff_eq_one, covby_iff_exists_erase, eq_comm] /-- `t` is in the shadow of `𝒜` iff we can add an element to it so that the resulting finset is in `𝒜`. -/ lemma mem_shadow_iff_insert_mem : t ∈ ∂ 𝒜 ↔ ∃ a, a ∉ t ∧ insert a t ∈ 𝒜 := by - simp_rw [mem_shadow_iff_exists_sdiff, ←covby_iff_card_sdiff_eq_one, covby_iff_exists_insert] + simp_rw [mem_shadow_iff_exists_sdiff, ← covby_iff_card_sdiff_eq_one, covby_iff_exists_insert] aesop #align finset.mem_shadow_iff_insert_mem Finset.mem_shadow_iff_insert_mem @@ -163,7 +163,7 @@ lemma _root_.Set.Sized.shadow_iterate (h𝒜 : (𝒜 : Set (Finset α)).Sized r) (∂^[k] 𝒜 : Set (Finset α)).Sized (r - k) := by simp_rw [Set.Sized, mem_coe, mem_shadow_iterate_iff_exists_sdiff] rintro t ⟨s, hs, hts, rfl⟩ - rw [card_sdiff hts, ←h𝒜 hs, Nat.sub_sub_self (card_le_of_subset hts)] + rw [card_sdiff hts, ← h𝒜 hs, Nat.sub_sub_self (card_le_of_subset hts)] theorem sized_shadow_iff (h : ∅ ∉ 𝒜) : (∂ 𝒜 : Set (Finset α)).Sized r ↔ (𝒜 : Set (Finset α)).Sized (r + 1) := by @@ -222,12 +222,12 @@ theorem insert_mem_upShadow (hs : s ∈ 𝒜) (ha : a ∉ s) : insert a s ∈ See also `Finset.mem_upShadow_iff_exists_mem_card_add_one`. -/ lemma mem_upShadow_iff_exists_sdiff : t ∈ ∂⁺ 𝒜 ↔ ∃ s ∈ 𝒜, s ⊆ t ∧ (t \ s).card = 1 := by - simp_rw [mem_upShadow_iff, ←covby_iff_card_sdiff_eq_one, covby_iff_exists_insert] + simp_rw [mem_upShadow_iff, ← covby_iff_card_sdiff_eq_one, covby_iff_exists_insert] /-- `t` is in the upper shadow of `𝒜` iff we can remove an element from it so that the resulting finset is in `𝒜`. -/ lemma mem_upShadow_iff_erase_mem : t ∈ ∂⁺ 𝒜 ↔ ∃ a, a ∈ t ∧ erase t a ∈ 𝒜 := by - simp_rw [mem_upShadow_iff_exists_sdiff, ←covby_iff_card_sdiff_eq_one, covby_iff_exists_erase] + simp_rw [mem_upShadow_iff_exists_sdiff, ← covby_iff_card_sdiff_eq_one, covby_iff_exists_erase] aesop #align finset.mem_up_shadow_iff_erase_mem Finset.mem_upShadow_iff_erase_mem @@ -247,7 +247,7 @@ lemma mem_upShadow_iterate_iff_exists_card : induction' k with k ih generalizing t · simp simp only [mem_upShadow_iff_erase_mem, ih, Function.iterate_succ_apply', card_eq_succ, - subset_erase, erase_sdiff_comm, ←sdiff_insert] + subset_erase, erase_sdiff_comm, ← sdiff_insert] constructor · rintro ⟨a, hat, u, rfl, ⟨hut, hau⟩, htu⟩ exact ⟨_, ⟨_, _, hau, rfl, rfl⟩, insert_subset hat hut, htu⟩ @@ -324,14 +324,14 @@ theorem mem_upShadow_iff_exists_mem_card_add : ext s simp only [mem_image, exists_prop, mem_shadow_iff, mem_upShadow_iff, mem_compls] refine (compl_involutive.toPerm _).exists_congr_left.trans ?_ - simp [←compl_involutive.eq_iff] + simp [← compl_involutive.eq_iff] #align finset.up_shadow_image_compl Finset.shadow_compls @[simp] lemma upShadow_compls : ∂⁺ 𝒜ᶜˢ = (∂ 𝒜)ᶜˢ := by ext s simp only [mem_image, exists_prop, mem_shadow_iff, mem_upShadow_iff, mem_compls] refine (compl_involutive.toPerm _).exists_congr_left.trans ?_ - simp [←compl_involutive.eq_iff] + simp [← compl_involutive.eq_iff] #align finset.shadow_image_compl Finset.upShadow_compls end UpShadow diff --git a/Mathlib/Combinatorics/SetFamily/Shatter.lean b/Mathlib/Combinatorics/SetFamily/Shatter.lean index 739051a44c187..f2ad695c504a8 100644 --- a/Mathlib/Combinatorics/SetFamily/Shatter.lean +++ b/Mathlib/Combinatorics/SetFamily/Shatter.lean @@ -56,14 +56,14 @@ protected lemma Shatters.nonempty (h : 𝒜.Shatters s) : 𝒜.Nonempty := let ⟨t, ht, _⟩ := h Subset.rfl; ⟨t, ht⟩ @[simp] lemma shatters_empty : 𝒜.Shatters ∅ ↔ 𝒜.Nonempty := - ⟨Shatters.nonempty, fun ⟨s, hs⟩ t ht ↦ ⟨s, hs, by rwa [empty_inter, eq_comm, ←subset_empty]⟩⟩ + ⟨Shatters.nonempty, fun ⟨s, hs⟩ t ht ↦ ⟨s, hs, by rwa [empty_inter, eq_comm, ← subset_empty]⟩⟩ protected lemma Shatters.subset_iff (h : 𝒜.Shatters s) : t ⊆ s ↔ ∃ u ∈ 𝒜, s ∩ u = t := ⟨fun ht ↦ h ht, by rintro ⟨u, _, rfl⟩; exact inter_subset_left _ _⟩ lemma shatters_iff : 𝒜.Shatters s ↔ 𝒜.image (fun t ↦ s ∩ t) = s.powerset := ⟨fun h ↦ by ext t; rw [mem_image, mem_powerset, h.subset_iff], - fun h t ht ↦ by rwa [←mem_powerset, ←h, mem_image] at ht⟩ + fun h t ht ↦ by rwa [← mem_powerset, ← h, mem_image] at ht⟩ lemma univ_shatters [Fintype α] : univ.Shatters s := shatters_of_forall_subset <| fun _ _ ↦ mem_univ _ @@ -90,7 +90,7 @@ lemma subset_shatterer (h : IsLowerSet (𝒜 : Set (Finset α))) : 𝒜 ⊆ 𝒜 @[simp] lemma shatterer_eq : 𝒜.shatterer = 𝒜 ↔ IsLowerSet (𝒜 : Set (Finset α)) := by refine ⟨fun h ↦ ?_, fun h ↦ Subset.antisymm (fun s hs ↦ ?_) <| subset_shatterer h⟩ - · rw [←h] + · rw [← h] exact isLowerSet_shatterer _ · obtain ⟨t, ht, hst⟩ := (mem_shatterer.1 hs).exists_superset exact h hst ht @@ -98,7 +98,7 @@ lemma subset_shatterer (h : IsLowerSet (𝒜 : Set (Finset α))) : 𝒜 ⊆ 𝒜 @[simp] lemma shatterer_idem : 𝒜.shatterer.shatterer = 𝒜.shatterer := by simp @[simp] lemma shatters_shatterer : 𝒜.shatterer.Shatters s ↔ 𝒜.Shatters s := by - simp_rw [←mem_shatterer, shatterer_idem] + simp_rw [← mem_shatterer, shatterer_idem] protected alias ⟨_, Shatters.shatterer⟩ := shatters_shatterer @@ -119,9 +119,9 @@ lemma card_le_card_shatterer (𝒜 : Finset (Finset α)) : 𝒜.card ≤ 𝒜.sh simp only [coe_inter, Set.subset_def, Set.mem_inter_iff, mem_coe, Set.mem_setOf_eq, and_imp, mem_shatterer] exact fun s _ ↦ aux (fun t ht ↦ (mem_filter.1 ht).2) - rw [←card_memberSubfamily_add_card_nonMemberSubfamily a] + rw [← card_memberSubfamily_add_card_nonMemberSubfamily a] refine (add_le_add ih₁ ih₀).trans ?_ - rw [←card_union_add_card_inter, ←hℬ, ←card_disjoint_union] + rw [← card_union_add_card_inter, ← hℬ, ← card_disjoint_union] swap · simp only [disjoint_left, mem_union, mem_shatterer, mem_image, not_exists, not_and] rintro _ (hs | hs) s - rfl @@ -135,7 +135,7 @@ lemma card_le_card_shatterer (𝒜 : Finset (Finset α)) : 𝒜.card ≤ 𝒜.sh refine ⟨insert a u, hu.1, inter_insert_of_not_mem fun ha ↦ ?_⟩ obtain ⟨v, hv, hsv⟩ := hs.exists_inter_eq_singleton ha rw [mem_memberSubfamily] at hv - rw [←singleton_subset_iff (a := a), ←hsv] at hv + rw [← singleton_subset_iff (a := a), ← hsv] at hv exact hv.2 <| inter_subset_right _ _ · refine forall_image.2 fun s hs ↦ mem_shatterer.2 fun t ht ↦ ?_ simp only [mem_inter, mem_shatterer] at hs @@ -144,7 +144,7 @@ lemma card_le_card_shatterer (𝒜 : Finset (Finset α)) : 𝒜.card ≤ 𝒜.sh · obtain ⟨u, hu, hsu⟩ := hs.1 ht rw [mem_memberSubfamily] at hu refine ⟨_, hu.1, ?_⟩ - rw [←insert_inter_distrib, hsu, insert_erase ha] + rw [← insert_inter_distrib, hsu, insert_erase ha] · obtain ⟨u, hu, hsu⟩ := hs.2 ht rw [mem_nonMemberSubfamily] at hu refine ⟨_, hu.1, ?_⟩ @@ -190,7 +190,7 @@ lemma vcDim_compress_le (a : α) (𝒜 : Finset (Finset α)) : (𝓓 a 𝒜).vcD /-- The **Sauer-Shelah lemma**. -/ lemma card_shatterer_le_sum_vcDim [Fintype α] : 𝒜.shatterer.card ≤ ∑ k in Iic 𝒜.vcDim, (Fintype.card α).choose k := by - simp_rw [←card_univ, ←card_powersetCard] + simp_rw [← card_univ, ← card_powersetCard] refine (card_le_of_subset <| fun s hs ↦ mem_biUnion.2 ⟨card s, ?_⟩).trans card_biUnion_le exact ⟨mem_Iic.2 (mem_shatterer.1 hs).card_le_vcDim, mem_powersetCard_univ.2 rfl⟩ diff --git a/Mathlib/Combinatorics/SimpleGraph/Basic.lean b/Mathlib/Combinatorics/SimpleGraph/Basic.lean index b254a8b6a7d54..14761348a2efa 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Basic.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Aaron Anderson, Jalex Stark, Kyle Miller, Alena Gusakov, Hunter Monroe -/ import Mathlib.Combinatorics.SimpleGraph.Init +import Mathlib.Data.FunLike.Fintype import Mathlib.Data.Rel import Mathlib.Data.Set.Finite import Mathlib.Data.Sym.Card @@ -420,6 +421,14 @@ theorem emptyGraph_eq_bot (V : Type u) : emptyGraph V = ⊥ := instance (V : Type u) : Inhabited (SimpleGraph V) := ⟨⊥⟩ +instance [Subsingleton V] : Unique (SimpleGraph V) where + default := ⊥ + uniq G := by ext a b; have := Subsingleton.elim a b; simp [this] + +instance [Nontrivial V] : Nontrivial (SimpleGraph V) := + ⟨⟨⊥, ⊤, fun h ↦ not_subsingleton V ⟨by simpa only [←adj_inj, Function.funext_iff, bot_adj, + top_adj, ne_eq, eq_iff_iff, false_iff, not_not] using h⟩⟩⟩ + section Decidable variable (V) (H : SimpleGraph V) [DecidableRel G.Adj] [DecidableRel H.Adj] @@ -565,6 +574,17 @@ theorem edgeSet_sdiff : (G₁ \ G₂).edgeSet = G₁.edgeSet \ G₂.edgeSet := b rfl #align simple_graph.edge_set_sdiff SimpleGraph.edgeSet_sdiff +variable {G G₁ G₂} + +@[simp] lemma disjoint_edgeSet : Disjoint G₁.edgeSet G₂.edgeSet ↔ Disjoint G₁ G₂ := by + rw [Set.disjoint_iff, disjoint_iff_inf_le, ←edgeSet_inf, ←edgeSet_bot, ←Set.le_iff_subset, + OrderEmbedding.le_iff_le] + +@[simp] lemma edgeSet_eq_empty : G.edgeSet = ∅ ↔ G = ⊥ := by rw [←edgeSet_bot, edgeSet_inj] + +@[simp] lemma edgeSet_nonempty : G.edgeSet.Nonempty ↔ G ≠ ⊥ := by + rw [Set.nonempty_iff_ne_empty, edgeSet_eq_empty.ne] + /-- This lemma, combined with `edgeSet_sdiff` and `edgeSet_from_edgeSet`, allows proving `(G \ from_edgeSet s).edge_set = G.edgeSet \ s` by `simp`. -/ @[simp] @@ -593,6 +613,8 @@ theorem adj_iff_exists_edge_coe : G.Adj a b ↔ ∃ e : G.edgeSet, e.val = ⟦(a simp only [mem_edgeSet, exists_prop, SetCoe.exists, exists_eq_right, Subtype.coe_mk] #align simple_graph.adj_iff_exists_edge_coe SimpleGraph.adj_iff_exists_edge_coe +variable (G G₁ G₂) + theorem edge_other_ne {e : Sym2 V} (he : e ∈ G.edgeSet) {v : V} (h : v ∈ e) : Sym2.Mem.other h ≠ v := by erw [← Sym2.other_spec h, Sym2.eq_swap] at he @@ -703,6 +725,14 @@ theorem fromEdgeSet_mono {s t : Set (Sym2 V)} (h : s ⊆ t) : fromEdgeSet s ≤ exact fun vws _ => h vws #align simple_graph.from_edge_set_mono SimpleGraph.fromEdgeSet_mono +@[simp] lemma disjoint_fromEdgeSet : Disjoint G (fromEdgeSet s) ↔ Disjoint G.edgeSet s := by + conv_rhs => rw [←Set.diff_union_inter s {e : Sym2 V | e.IsDiag}] + rw [←disjoint_edgeSet, edgeSet_fromEdgeSet, Set.disjoint_union_right, and_iff_left] + exact Set.disjoint_left.2 fun e he he' ↦ not_isDiag_of_mem_edgeSet _ he he'.2 + +@[simp] lemma fromEdgeSet_disjoint : Disjoint (fromEdgeSet s) G ↔ Disjoint s G.edgeSet := by + rw [disjoint_comm, disjoint_fromEdgeSet, disjoint_comm] + instance [DecidableEq V] [Fintype s] : Fintype (fromEdgeSet s).edgeSet := by rw [edgeSet_fromEdgeSet s] infer_instance @@ -1167,6 +1197,9 @@ theorem deleteEdges_eq_sdiff_fromEdgeSet (s : Set (Sym2 V)) : exact ⟨fun h => ⟨h.1, not_and_of_not_left _ h.2⟩, fun h => ⟨h.1, not_and'.mp h.2 h.ne⟩⟩ #align simple_graph.delete_edges_eq_sdiff_from_edge_set SimpleGraph.deleteEdges_eq_sdiff_fromEdgeSet +@[simp] lemma deleteEdges_eq {s : Set (Sym2 V)} : G.deleteEdges s = G ↔ Disjoint G.edgeSet s := by + rw [deleteEdges_eq_sdiff_fromEdgeSet, sdiff_eq_left, disjoint_fromEdgeSet] + theorem compl_eq_deleteEdges : Gᶜ = (⊤ : SimpleGraph V).deleteEdges G.edgeSet := by ext simp @@ -1241,18 +1274,18 @@ def DeleteFar (p : SimpleGraph V → Prop) (r : 𝕜) : Prop := ∀ ⦃s⦄, s ⊆ G.edgeFinset → p (G.deleteEdges s) → r ≤ s.card #align simple_graph.delete_far SimpleGraph.DeleteFar -open Classical - variable {G} theorem deleteFar_iff : - G.DeleteFar p r ↔ ∀ ⦃H⦄, H ≤ G → p H → r ≤ G.edgeFinset.card - H.edgeFinset.card := by - refine' ⟨fun h H hHG hH => _, fun h s hs hG => _⟩ + G.DeleteFar p r ↔ ∀ ⦃H : SimpleGraph _⦄ [DecidableRel H.Adj], + H ≤ G → p H → r ≤ G.edgeFinset.card - H.edgeFinset.card := by + refine' ⟨fun h H _ hHG hH => _, fun h s hs hG => _⟩ · have := h (sdiff_subset G.edgeFinset H.edgeFinset) simp only [deleteEdges_sdiff_eq_of_le _ hHG, edgeFinset_mono hHG, card_sdiff, card_le_of_subset, coe_sdiff, coe_edgeFinset, Nat.cast_sub] at this - exact this hH - · simpa [card_sdiff hs, edgeFinset_deleteEdges, -Set.toFinset_card, Nat.cast_sub, + convert this hH + · classical + simpa [card_sdiff hs, edgeFinset_deleteEdges, -Set.toFinset_card, Nat.cast_sub, card_le_of_subset hs] using h (G.deleteEdges_le s) hG #align simple_graph.delete_far_iff SimpleGraph.deleteFar_iff @@ -1325,23 +1358,53 @@ theorem map_adj (f : V ↪ W) (G : SimpleGraph V) (u v : W) : Iff.rfl #align simple_graph.map_adj SimpleGraph.map_adj +lemma map_adj_apply {G : SimpleGraph V} {f : V ↪ W} {a b : V} : + (G.map f).Adj (f a) (f b) ↔ G.Adj a b := by simp + theorem map_monotone (f : V ↪ W) : Monotone (SimpleGraph.map f) := by rintro G G' h _ _ ⟨u, v, ha, rfl, rfl⟩ exact ⟨_, _, h ha, rfl, rfl⟩ #align simple_graph.map_monotone SimpleGraph.map_monotone +@[simp] lemma map_id : G.map (Function.Embedding.refl _) = G := + SimpleGraph.ext _ _ $ Relation.map_id_id _ + +@[simp] lemma map_map (f : V ↪ W) (g : W ↪ X) : (G.map f).map g = G.map (f.trans g) := + SimpleGraph.ext _ _ $ Relation.map_map _ _ _ _ _ + +instance instDecidableMapAdj (f : V ↪ W) (G : SimpleGraph V) + [DecidableRel (Relation.Map G.Adj f f)] : DecidableRel (G.map f).Adj := ‹DecidableRel _› + /-- Given a function, there is a contravariant induced map on graphs by pulling back the adjacency relation. This is one of the ways of creating induced graphs. See `SimpleGraph.induce` for a wrapper. This is surjective when `f` is injective (see `SimpleGraph.comap_surjective`).-/ -@[simps] protected def comap (f : V → W) (G : SimpleGraph W) : SimpleGraph V where Adj u v := G.Adj (f u) (f v) symm _ _ h := h.symm loopless _ := G.loopless _ #align simple_graph.comap SimpleGraph.comap +@[simp] lemma comap_adj {G : SimpleGraph W} {f : V → W} : + (G.comap f).Adj u v ↔ G.Adj (f u) (f v) := Iff.rfl + +@[simp] lemma comap_id {G : SimpleGraph V} : G.comap id = G := SimpleGraph.ext _ _ rfl + +@[simp] lemma comap_comap {G : SimpleGraph X} (f : V → W) (g : W → X) : + (G.comap g).comap f = G.comap (g ∘ f) := rfl + +instance instDecidableComapAdj (f : V → W) (G : SimpleGraph W) [DecidableRel G.Adj] : + DecidableRel (G.comap f).Adj := fun _ _ ↦ ‹DecidableRel G.Adj› _ _ + +lemma comap_symm (G : SimpleGraph V) (e : V ≃ W) : + G.comap e.symm.toEmbedding = G.map e.toEmbedding := by + ext; simp only [Equiv.apply_eq_iff_eq_symm_apply, comap_adj, map_adj, Equiv.toEmbedding_apply, + exists_eq_right_right, exists_eq_right] + +lemma map_symm (G : SimpleGraph W) (e : V ≃ W) : + G.map e.symm.toEmbedding = G.comap e.toEmbedding := by rw [←comap_symm, e.symm_symm] + theorem comap_monotone (f : V ↪ W) : Monotone (SimpleGraph.comap f) := by intro G G' h _ _ ha exact h ha @@ -1389,6 +1452,23 @@ Two vertices are adjacent if and only if their indices are not equal. -/ abbrev completeMultipartiteGraph {ι : Type*} (V : ι → Type*) : SimpleGraph (Σ i, V i) := SimpleGraph.comap Sigma.fst ⊤ +/-- Equivalent types have equivalent simple graphs. -/ +@[simps apply] +protected def _root_.Equiv.simpleGraph (e : V ≃ W) : SimpleGraph V ≃ SimpleGraph W where + toFun := SimpleGraph.comap e.symm + invFun := SimpleGraph.comap e + left_inv _ := by simp + right_inv _ := by simp + +@[simp] lemma _root_.Equiv.simpleGraph_refl : (Equiv.refl V).simpleGraph = Equiv.refl _ := by + ext; rfl + +@[simp] lemma _root_.Equiv.simpleGraph_trans (e₁ : V ≃ W) (e₂ : W ≃ X) : + (e₁.trans e₂).simpleGraph = e₁.simpleGraph.trans e₂.simpleGraph := rfl + +@[simp] +lemma _root_.Equiv.symm_simpleGraph (e : V ≃ W) : e.simpleGraph.symm = e.symm.simpleGraph := rfl + /-! ## Induced graphs -/ /- Given a set `s` of vertices, we can restrict a graph to those vertices by restricting its @@ -1771,13 +1851,30 @@ infixl:50 " ≃g " => Iso namespace Hom -variable {G G'} (f : G →g G') +variable {G G'} {G₁ G₂ : SimpleGraph V} {H : SimpleGraph W} (f : G →g G') /-- The identity homomorphism from a graph to itself. -/ -abbrev id : G →g G := +protected abbrev id : G →g G := RelHom.id _ #align simple_graph.hom.id SimpleGraph.Hom.id +@[simp, norm_cast] lemma coe_id : ⇑(Hom.id : G →g G) = _root_.id := rfl + +instance [Subsingleton (V → W)] : Subsingleton (G →g H) := + FunLike.coe_injective.subsingleton + +instance [IsEmpty V] : Unique (G →g H) where + default := ⟨isEmptyElim, fun {a} ↦ isEmptyElim a⟩ + uniq _ := Subsingleton.elim _ _ + +instance instFintype [DecidableEq V] [Fintype V] [Fintype W] [DecidableRel G.Adj] + [DecidableRel H.Adj] : Fintype (G →g H) := + Fintype.ofEquiv {f : V → W // ∀ {a b}, G.Adj a b → H.Adj (f a) (f b)} + { toFun := fun f ↦ ⟨f.1, f.2⟩, invFun := fun f ↦ ⟨f.1, f.2⟩, + left_inv := fun _ ↦ rfl, right_inv := fun _ ↦ rfl } + +instance [Finite V] [Finite W] : Finite (G →g H) := FunLike.finite _ + theorem map_adj {v w : V} (h : G.Adj v w) : G'.Adj (f v) (f w) := f.map_rel' h #align simple_graph.hom.map_adj SimpleGraph.Hom.map_adj @@ -1854,11 +1951,16 @@ theorem coe_comp (f' : G' →g G'') (f : G →g G') : ⇑(f'.comp f) = f' ∘ f rfl #align simple_graph.hom.coe_comp SimpleGraph.Hom.coe_comp +/-- The graph homomorphism from a smaller graph to a bigger one. -/ +def ofLe (h : G₁ ≤ G₂) : G₁ →g G₂ := ⟨id, @h⟩ + +@[simp, norm_cast] lemma coe_ofLe (h : G₁ ≤ G₂) : ⇑(ofLe h) = id := rfl + end Hom namespace Embedding -variable {G G'} (f : G ↪g G') +variable {G G'} {H : SimpleGraph W} (f : G ↪g G') /-- The identity embedding from a graph to itself. -/ abbrev refl : G ↪g G := @@ -1870,7 +1972,9 @@ abbrev toHom : G →g G' := f.toRelHom #align simple_graph.embedding.to_hom SimpleGraph.Embedding.toHom -theorem map_adj_iff {v w : V} : G'.Adj (f v) (f w) ↔ G.Adj v w := +@[simp] lemma coe_toHom (f : G ↪g H) : ⇑f.toHom = f := rfl + +@[simp] theorem map_adj_iff {v w : V} : G'.Adj (f v) (f w) ↔ G.Adj v w := f.map_rel_iff #align simple_graph.embedding.map_adj_iff SimpleGraph.Embedding.map_adj_iff diff --git a/Mathlib/Combinatorics/SimpleGraph/Connectivity.lean b/Mathlib/Combinatorics/SimpleGraph/Connectivity.lean index dae3cd0fa64e3..671f4ac3d2cd1 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Connectivity.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Connectivity.lean @@ -2623,7 +2623,7 @@ theorem adj_and_reachable_delete_edges_iff_exists_cycle {v w : V} : · simp only [Sym2.eq_swap, Walk.edges_cons, List.mem_cons, eq_self_iff_true, true_or_iff] · rintro ⟨u, c, hc, he⟩ refine ⟨c.adj_of_mem_edges he, ?_⟩ - by_contra' hb + by_contra! hb have hb' : ∀ p : G.Walk w v, ⟦(w, v)⟧ ∈ p.edges := by intro p simpa [Sym2.eq_swap] using hb p.reverse diff --git a/Mathlib/Combinatorics/SimpleGraph/Ends/Defs.lean b/Mathlib/Combinatorics/SimpleGraph/Ends/Defs.lean index a29d6f6d07630..b0bb6263e91f4 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Ends/Defs.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Ends/Defs.lean @@ -159,7 +159,7 @@ theorem exists_adj_boundary_pair (Gc : G.Preconnected) (hK : K.Nonempty) : refine' ComponentCompl.ind fun v vnK => _ let C : G.ComponentCompl K := G.componentComplMk vnK let dis := Set.disjoint_iff.mp C.disjoint_right - by_contra' h + by_contra! h suffices Set.univ = (C : Set V) by exact dis ⟨hK.choose_spec, this ▸ Set.mem_univ hK.some⟩ symm rw [Set.eq_univ_iff_forall] diff --git a/Mathlib/Combinatorics/SimpleGraph/Girth.lean b/Mathlib/Combinatorics/SimpleGraph/Girth.lean index 158a0764292bd..8a15100c3bb14 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Girth.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Girth.lean @@ -36,8 +36,8 @@ lemma exists_girth_eq_length : refine' ⟨_, λ h ↦ _⟩ · rintro ⟨a, w, hw, _⟩ hG exact hG _ hw - · simp_rw [←girth_eq_top, ←Ne.def, girth, iInf_subtype', iInf_sigma', ENat.iInf_coe_ne_top, - ←exists_prop, Subtype.exists', Sigma.exists', eq_comm] at h ⊢ + · simp_rw [← girth_eq_top, ← Ne.def, girth, iInf_subtype', iInf_sigma', ENat.iInf_coe_ne_top, + ← exists_prop, Subtype.exists', Sigma.exists', eq_comm] at h ⊢ exact ciInf_mem _ @[simp] lemma girth_bot : girth (⊥ : SimpleGraph α) = ⊤ := by simp diff --git a/Mathlib/Combinatorics/SimpleGraph/Regularity/Bound.lean b/Mathlib/Combinatorics/SimpleGraph/Regularity/Bound.lean index 4fe57467013c8..866a22b248f8c 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Regularity/Bound.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Regularity/Bound.lean @@ -251,8 +251,8 @@ theorem add_div_le_sum_sq_div_card (hst : s ⊆ t) (f : ι → 𝕜) (d : 𝕜) apply (add_le_add_left h₃ _).trans -- Porting note: was -- `simp [← mul_div_right_comm _ (t.card : 𝕜), sub_div' _ _ _ htcard.ne', ← sum_div, ← add_div,` - -- ` mul_pow, div_le_iff (sq_pos_of_ne_zero _ htcard.ne'), sub_sq, sum_add_distrib, ← sum_mul, ←` - -- ` mul_sum]` + -- ` mul_pow, div_le_iff (sq_pos_of_ne_zero _ htcard.ne'), sub_sq, sum_add_distrib, ← sum_mul,` + -- ` ← mul_sum]` simp_rw [sub_div' _ _ _ htcard.ne'] conv_lhs => enter [2, 2, x]; rw [div_pow] rw [div_pow, ← sum_div, ← mul_div_right_comm _ (t.card : 𝕜), ← add_div, diff --git a/Mathlib/Combinatorics/SimpleGraph/Regularity/Uniform.lean b/Mathlib/Combinatorics/SimpleGraph/Regularity/Uniform.lean index 13032102df34a..e16868c4389f3 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Regularity/Uniform.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Regularity/Uniform.lean @@ -220,7 +220,7 @@ theorem nonUniforms_bot (hε : 0 < ε) : (⊥ : Finpartition A).nonUniforms G ε simp only [Finpartition.mk_mem_nonUniforms_iff, Finpartition.parts_bot, mem_map, not_and, Classical.not_not, exists_imp]; dsimp rintro x ⟨_,xu⟩ y ⟨_,yv⟩ _ - rw [←xu, ←yv] + rw [← xu, ← yv] exact G.isUniform_singleton hε #align finpartition.non_uniforms_bot Finpartition.nonUniforms_bot diff --git a/Mathlib/Combinatorics/SimpleGraph/Triangle/Basic.lean b/Mathlib/Combinatorics/SimpleGraph/Triangle/Basic.lean index 1293321f8ce43..bb7ad8a1b8f87 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Triangle/Basic.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Triangle/Basic.lean @@ -42,8 +42,9 @@ def FarFromTriangleFree (G : SimpleGraph α) (ε : 𝕜) : Prop := (G.DeleteFar fun H => H.CliqueFree 3) <| ε * (card α ^ 2 : ℕ) #align simple_graph.far_from_triangle_free SimpleGraph.FarFromTriangleFree -theorem farFromTriangleFree_iff : G.FarFromTriangleFree ε ↔ ∀ ⦃H⦄, H ≤ G → H.CliqueFree 3 → - ε * (card α ^ 2 : ℕ) ≤ G.edgeFinset.card - H.edgeFinset.card := deleteFar_iff +theorem farFromTriangleFree_iff : + G.FarFromTriangleFree ε ↔ ∀ ⦃H : SimpleGraph α⦄, [DecidableRel H.Adj] → H ≤ G → H.CliqueFree 3 → + ε * (card α ^ 2 : ℕ) ≤ G.edgeFinset.card - H.edgeFinset.card := deleteFar_iff #align simple_graph.far_from_triangle_free_iff SimpleGraph.farFromTriangleFree_iff alias ⟨farFromTriangleFree.le_card_sub_card, _⟩ := farFromTriangleFree_iff diff --git a/Mathlib/Combinatorics/Young/YoungDiagram.lean b/Mathlib/Combinatorics/Young/YoungDiagram.lean index ec26bfaaa6eda..e226d162bdd15 100644 --- a/Mathlib/Combinatorics/Young/YoungDiagram.lean +++ b/Mathlib/Combinatorics/Young/YoungDiagram.lean @@ -247,7 +247,7 @@ theorem transpose_le_iff {μ ν : YoungDiagram} : μ.transpose ≤ ν.transpose ⟨fun h => by convert YoungDiagram.le_of_transpose_le h simp, fun h => by - rw [←transpose_transpose μ] at h + rw [← transpose_transpose μ] at h exact YoungDiagram.le_of_transpose_le h ⟩ #align young_diagram.transpose_le_iff YoungDiagram.transpose_le_iff @@ -325,7 +325,7 @@ theorem rowLen_eq_card (μ : YoungDiagram) {i : ℕ} : μ.rowLen i = (μ.row i). @[mono] theorem rowLen_anti (μ : YoungDiagram) (i1 i2 : ℕ) (hi : i1 ≤ i2) : μ.rowLen i2 ≤ μ.rowLen i1 := by - by_contra' h_lt + by_contra! h_lt rw [← lt_self_iff_false (μ.rowLen i1)] rw [← mem_iff_lt_rowLen] at h_lt ⊢ exact μ.up_left_mem hi (by rfl) h_lt diff --git a/Mathlib/Computability/AkraBazzi/GrowsPolynomially.lean b/Mathlib/Computability/AkraBazzi/GrowsPolynomially.lean index 24fd536352dc4..889bd9b1aa119 100644 --- a/Mathlib/Computability/AkraBazzi/GrowsPolynomially.lean +++ b/Mathlib/Computability/AkraBazzi/GrowsPolynomially.lean @@ -647,11 +647,11 @@ lemma growsPolynomially_log : GrowsPolynomially Real.log := by 1 / 2 * Real.log x = Real.log x + (-1 / 2) * Real.log x := by ring _ ≤ Real.log x + Real.log b := by gcongr - rw [neg_div, neg_mul, ←neg_le] + rw [neg_div, neg_mul, ← neg_le] refine le_of_lt (hx x ?_) calc b * x ≤ 1 * x := by gcongr; exact le_of_lt hb.2 _ = x := by rw [one_mul] - _ = Real.log (b * x) := by rw [←Real.log_mul (by positivity) (by positivity), mul_comm] + _ = Real.log (b * x) := by rw [← Real.log_mul (by positivity) (by positivity), mul_comm] _ ≤ Real.log u := by gcongr; exact hu.1 case ub => rw [one_mul] @@ -691,21 +691,21 @@ lemma GrowsPolynomially.of_isTheta {f g : ℝ → ℝ} (hg : GrowsPolynomially g refine ⟨?lb, ?ub⟩ case lb => calc c₁ * c₂⁻¹ * c₃ * f x ≤ c₁ * c₂⁻¹ * c₃ * (c₂ * ‖g x‖) := by - rw [←Real.norm_of_nonneg (hf_pos x hbx)]; gcongr; exact h_ub x hbx + rw [← Real.norm_of_nonneg (hf_pos x hbx)]; gcongr; exact h_ub x hbx _ = (c₂⁻¹ * c₂) * c₁ * (c₃ * ‖g x‖) := by ring _ = c₁ * (c₃ * ‖g x‖) := by simp [c₂_cancel] _ ≤ c₁ * ‖g u‖ := by gcongr; exact (hg_bound u hu).1 _ ≤ f u := by - rw [←Real.norm_of_nonneg (hf_pos u hu.1)] + rw [← Real.norm_of_nonneg (hf_pos u hu.1)] exact h_lb u hu.1 case ub => calc - f u ≤ c₂ * ‖g u‖ := by rw [←Real.norm_of_nonneg (hf_pos u hu.1)]; exact h_ub u hu.1 + f u ≤ c₂ * ‖g u‖ := by rw [← Real.norm_of_nonneg (hf_pos u hu.1)]; exact h_ub u hu.1 _ ≤ c₂ * (c₄ * ‖g x‖) := by gcongr; exact (hg_bound u hu).2 _ = c₂ * c₄ * (c₁⁻¹ * c₁) * ‖g x‖ := by simp [c₁_cancel]; ring _ = c₂ * c₄ * c₁⁻¹ * (c₁ * ‖g x‖) := by ring _ ≤ c₂ * c₄ * c₁⁻¹ * f x := by gcongr - rw [←Real.norm_of_nonneg (hf_pos x hbx)] + rw [← Real.norm_of_nonneg (hf_pos x hbx)] exact h_lb x hbx lemma GrowsPolynomially.of_isEquivalent {f g : ℝ → ℝ} (hg : GrowsPolynomially g) diff --git a/Mathlib/Computability/Language.lean b/Mathlib/Computability/Language.lean index dacbc12e7e79a..d61641681a51a 100644 --- a/Mathlib/Computability/Language.lean +++ b/Mathlib/Computability/Language.lean @@ -294,7 +294,7 @@ instance : KleeneAlgebra (Language α) := refine' iSup_le (fun n ↦ _) induction' n with n ih · simp - rw [pow_succ, ←mul_assoc m l (l^n)] + rw [pow_succ, ← mul_assoc m l (l^n)] exact le_trans (le_mul_congr h le_rfl) ih } end Language diff --git a/Mathlib/Control/Random.lean b/Mathlib/Control/Random.lean index b09710502ceb9..4464d65c97046 100644 --- a/Mathlib/Control/Random.lean +++ b/Mathlib/Control/Random.lean @@ -54,21 +54,21 @@ class BoundedRandom (α : Type u) [Preorder α] where namespace Rand /-- Generate one more `Nat` -/ def next [RandomGen g] : RandG g Nat := do - let rng := (←get).down + let rng := (← get).down let (res, new) := RandomGen.next rng set (ULift.up new) pure res /-- Create a new random number generator distinct from the one stored in the state -/ def split {g : Type} [RandomGen g] : RandG g g := do - let rng := (←get).down + let rng := (← get).down let (r1, r2) := RandomGen.split rng set (ULift.up r1) pure r2 /-- Get the range of Nat that can be generated by the generator `g` -/ def range {g : Type} [RandomGen g] : RandG g (Nat × Nat) := do - let rng := (←get).down + let rng := (← get).down pure <| RandomGen.range rng end Rand diff --git a/Mathlib/Data/BitVec/Defs.lean b/Mathlib/Data/BitVec/Defs.lean index ae6f555ce378b..d42b5be99a64e 100644 --- a/Mathlib/Data/BitVec/Defs.lean +++ b/Mathlib/Data/BitVec/Defs.lean @@ -58,9 +58,11 @@ namespace Std.BitVec /-! ### Arithmetic operators -/ #align bitvec.neg Std.BitVec.neg -/-- Add with carry (no overflow) -/ -def adc {n} (x y : BitVec n) (c : Bool) : BitVec (n+1) := - ofFin (x.toNat + y.toNat + c.toNat) +/-- Add with carry (no overflow) + +See also `Std.BitVec.adc`, which stores the carry bit separately. -/ +def adc' {n} (x y : BitVec n) (c : Bool) : BitVec (n+1) := + let a := x.adc y c; .cons a.1 a.2 #align bitvec.adc Std.BitVec.adc #align bitvec.add Std.BitVec.add diff --git a/Mathlib/Data/BitVec/Lemmas.lean b/Mathlib/Data/BitVec/Lemmas.lean index ec8ffe729035e..97f2353f1d29f 100644 --- a/Mathlib/Data/BitVec/Lemmas.lean +++ b/Mathlib/Data/BitVec/Lemmas.lean @@ -31,20 +31,10 @@ theorem toNat_inj {x y : BitVec w} : x.toNat = y.toNat ↔ x = y := theorem toNat_lt_toNat {x y : BitVec w} : x.toNat < y.toNat ↔ x < y := Iff.rfl -@[simp] -lemma ofNat_eq_mod_two_pow (n : Nat) : (BitVec.ofNat w n).toNat = n % 2^w := rfl - -lemma toNat_ofNat {m} (h : m < 2^w) : (BitVec.ofNat w m).toNat = m := Fin.val_cast_of_lt h - -@[simp] -lemma toNat_ofFin (x : Fin (2^w)) : (ofFin x).toNat = x.val := rfl +attribute [simp] toNat_ofNat toNat_ofFin -theorem toNat_append (msbs : BitVec w) (lsbs : BitVec v) : - (msbs ++ lsbs).toNat = msbs.toNat <<< v ||| lsbs.toNat := by - rcases msbs with ⟨msbs, hm⟩ - rcases lsbs with ⟨lsbs, hl⟩ - simp only [HAppend.hAppend, append, toNat_ofFin] - rw [toNat_ofNat (Nat.add_comm w v ▸ append_lt hl hm)] +lemma toNat_ofNat_of_lt {m} (h : m < 2^w) : (BitVec.ofNat w m).toNat = m := by + simp only [toNat_ofNat, mod_eq_of_lt h] #noalign bitvec.bits_to_nat_to_bool @@ -60,14 +50,9 @@ lemma extractLsb_eq {w : ℕ} (hi lo : ℕ) (a : BitVec w) : theorem toNat_extractLsb' {i j} {x : BitVec w} : (extractLsb' i j x).toNat = x.toNat / 2 ^ i % (2 ^ j) := by - simp only [extractLsb', ofNat_eq_mod_two_pow, shiftRight_eq_div_pow] - -theorem getLsb_eq_testBit {i} {x : BitVec w} : getLsb x i = x.toNat.testBit i := by - simp only [getLsb, Nat.shiftLeft_eq, one_mul, Nat.and_two_pow] - cases' testBit (BitVec.toNat x) i - <;> simp [pos_iff_ne_zero.mp (two_pow_pos i)] + simp only [extractLsb', toNat_ofNat, shiftRight_eq_div_pow] -theorem ofFin_val {n : ℕ} (i : Fin <| 2 ^ n) : (ofFin i).toNat = i.val := by +theorem ofFin_val {n : ℕ} (i : Fin <| 2 ^ n) : (ofFin i).toNat = i.val := rfl #align bitvec.of_fin_val Std.BitVec.ofFin_val @@ -91,16 +76,12 @@ theorem decide_addLsb_mod_two {x b} : decide (addLsb x b % 2 = 1) = b := by simp [addLsb] #align bitvec.to_bool_add_lsb_mod_two Std.BitVec.decide_addLsb_mod_two -@[simp] -lemma ofNat_toNat (x : BitVec w) : BitVec.ofNat w x.toNat = x := by - rcases x with ⟨x⟩ - simp [BitVec.ofNat] - apply Fin.cast_val_eq_self x -#align bitvec.of_nat_to_nat Std.BitVec.ofNat_toNat +lemma ofNat_toNat' (x : BitVec w) : (x.toNat)#w = x := by + rw [ofNat_toNat, truncate_eq] -lemma ofNat_toNat' (x : BitVec w) (h : w = v): +lemma ofNat_toNat_of_eq (x : BitVec w) (h : w = v): BitVec.ofNat v x.toNat = x.cast h := by - cases h; rw [ofNat_toNat, cast_eq] + cases h; rw [ofNat_toNat', cast_eq] theorem toFin_val {n : ℕ} (v : BitVec n) : (toFin v : ℕ) = v.toNat := by rfl diff --git a/Mathlib/Data/Bool/Basic.lean b/Mathlib/Data/Bool/Basic.lean index 451b1f5470010..50bddc60ebbe1 100644 --- a/Mathlib/Data/Bool/Basic.lean +++ b/Mathlib/Data/Bool/Basic.lean @@ -47,14 +47,7 @@ theorem of_decide_iff {p : Prop} [Decidable p] : decide p ↔ p := coe_decide p #align bool.of_to_bool_iff Bool.of_decide_iff -@[simp] -theorem true_eq_decide_iff {p : Prop} [Decidable p] : true = decide p ↔ p := - eq_comm.trans of_decide_iff #align bool.tt_eq_to_bool_iff Bool.true_eq_decide_iff - -@[simp] -theorem false_eq_decide_iff {p : Prop} [Decidable p] : false = decide p ↔ ¬p := - eq_comm.trans (decide_false_iff _) #align bool.ff_eq_to_bool_iff Bool.false_eq_decide_iff theorem decide_not (p : Prop) [Decidable p] : (decide ¬p) = !(decide p) := by @@ -308,22 +301,13 @@ theorem right_le_or : ∀ x y : Bool, y ≤ (x || y) := by decide theorem or_le : ∀ {x y z}, x ≤ z → y ≤ z → (x || y) ≤ z := by decide #align bool.bor_le Bool.or_le -/-- convert a `Bool` to a `ℕ`, `false -> 0`, `true -> 1` -/ -def toNat (b : Bool) : Nat := - cond b 1 0 #align bool.to_nat Bool.toNat -lemma toNat_le_one (b : Bool) : b.toNat ≤ 1 := by - cases b <;> decide - /-- convert a `ℕ` to a `Bool`, `0 -> false`, everything else -> `true` -/ def ofNat (n : Nat) : Bool := decide (n ≠ 0) #align bool.of_nat Bool.ofNat -@[simp] lemma toNat_true : toNat true = 1 := rfl -@[simp] lemma toNat_false : toNat false = 0 := rfl - @[simp] lemma toNat_beq_zero (b : Bool) : (b.toNat == 0) = !b := by cases b <;> rfl @[simp] lemma toNat_bne_zero (b : Bool) : (b.toNat != 0) = b := by simp [bne] @[simp] lemma toNat_beq_one (b : Bool) : (b.toNat == 1) = b := by cases b <;> rfl diff --git a/Mathlib/Data/Complex/BigOperators.lean b/Mathlib/Data/Complex/BigOperators.lean index 2773653161d13..6488381312002 100644 --- a/Mathlib/Data/Complex/BigOperators.lean +++ b/Mathlib/Data/Complex/BigOperators.lean @@ -31,12 +31,12 @@ theorem ofReal_sum (f : α → ℝ) : ((∑ i in s, f i : ℝ) : ℂ) = ∑ i in @[simp] theorem re_sum (f : α → ℂ) : (∑ i in s, f i).re = ∑ i in s, (f i).re := - reAddGroupHom.map_sum f s + map_sum reAddGroupHom f s #align complex.re_sum Complex.re_sum @[simp] theorem im_sum (f : α → ℂ) : (∑ i in s, f i).im = ∑ i in s, (f i).im := - imAddGroupHom.map_sum f s + map_sum imAddGroupHom f s #align complex.im_sum Complex.im_sum end Complex diff --git a/Mathlib/Data/Complex/Exponential.lean b/Mathlib/Data/Complex/Exponential.lean index 27616587f1ad9..e3ca019eeea19 100644 --- a/Mathlib/Data/Complex/Exponential.lean +++ b/Mathlib/Data/Complex/Exponential.lean @@ -142,7 +142,7 @@ theorem isCauSeq_geo_series {β : Type*} [Ring β] [Nontrivial β] {abv : β → refine' div_nonneg (sub_nonneg.2 _) (sub_nonneg.2 <| le_of_lt hx1) exact pow_le_one _ (by positivity) hx1.le · intro n _ - rw [←one_mul (abv x ^ n), pow_succ] + rw [← one_mul (abv x ^ n), pow_succ] gcongr) #align is_cau_geo_series isCauSeq_geo_series diff --git a/Mathlib/Data/Complex/Module.lean b/Mathlib/Data/Complex/Module.lean index 0db968ab8db87..3de9c46099c40 100644 --- a/Mathlib/Data/Complex/Module.lean +++ b/Mathlib/Data/Complex/Module.lean @@ -461,7 +461,7 @@ theorem imaginaryPart_I_smul (a : A) : ℑ (I • a) = ℜ a := by -- Porting note: was -- simp [smul_comm I, smul_smul I] rw [realPart_apply_coe, imaginaryPart_apply_coe, smul_comm] - simp [←smul_assoc] + simp [← smul_assoc] set_option linter.uppercaseLean3 false in #align imaginary_part_I_smul imaginaryPart_I_smul @@ -527,7 +527,7 @@ open Submodule lemma span_selfAdjoint : span ℂ (selfAdjoint A : Set A) = ⊤ := by refine eq_top_iff'.mpr fun x ↦ ?_ - rw [←realPart_add_I_smul_imaginaryPart x] + rw [← realPart_add_I_smul_imaginaryPart x] exact add_mem (subset_span (ℜ x).property) <| SMulMemClass.smul_mem _ <| subset_span (ℑ x).property @@ -548,7 +548,7 @@ lemma Complex.coe_selfAdjointEquiv (z : selfAdjoint ℂ) : @[simp] lemma realPart_ofReal (r : ℝ) : (ℜ (r : ℂ) : ℂ) = r := by - rw [realPart_apply_coe, star_def, conj_ofReal, ←two_smul ℝ (r : ℂ)] + rw [realPart_apply_coe, star_def, conj_ofReal, ← two_smul ℝ (r : ℂ)] simp @[simp] @@ -558,8 +558,8 @@ lemma imaginaryPart_ofReal (r : ℝ) : ℑ (r : ℂ) = 0 := by lemma Complex.coe_realPart (z : ℂ) : (ℜ z : ℂ) = z.re := calc (ℜ z : ℂ) = _ := by congrm (ℜ $((re_add_im z).symm)) _ = z.re := by - rw [map_add, AddSubmonoid.coe_add, mul_comm, ←smul_eq_mul, realPart_I_smul] - simp [conj_ofReal, ←two_mul] + rw [map_add, AddSubmonoid.coe_add, mul_comm, ← smul_eq_mul, realPart_I_smul] + simp [conj_ofReal, ← two_mul] end RealImaginaryPart section Rational diff --git a/Mathlib/Data/DFinsupp/Basic.lean b/Mathlib/Data/DFinsupp/Basic.lean index ff6e72fdbc4f2..c3c3eb5eb44f2 100644 --- a/Mathlib/Data/DFinsupp/Basic.lean +++ b/Mathlib/Data/DFinsupp/Basic.lean @@ -296,13 +296,13 @@ instance addCommMonoid [∀ i, AddCommMonoid (β i)] : AddCommMonoid (Π₀ i, @[simp] theorem coe_finset_sum {α} [∀ i, AddCommMonoid (β i)] (s : Finset α) (g : α → Π₀ i, β i) : ⇑(∑ a in s, g a) = ∑ a in s, ⇑(g a) := - (coeFnAddMonoidHom : _ →+ ∀ i, β i).map_sum g s + map_sum coeFnAddMonoidHom g s #align dfinsupp.coe_finset_sum DFinsupp.coe_finset_sum @[simp] theorem finset_sum_apply {α} [∀ i, AddCommMonoid (β i)] (s : Finset α) (g : α → Π₀ i, β i) (i : ι) : (∑ a in s, g a) i = ∑ a in s, g a i := - (evalAddMonoidHom i : _ →+ β i).map_sum g s + map_sum (evalAddMonoidHom i) g s #align dfinsupp.finset_sum_apply DFinsupp.finset_sum_apply instance [∀ i, AddGroup (β i)] : Neg (Π₀ i, β i) := @@ -1800,7 +1800,7 @@ theorem prod_comm {ι₁ ι₂ : Sort _} {β₁ : ι₁ → Type*} {β₂ : ι theorem sum_apply {ι₁ : Type u₁} [DecidableEq ι₁] {β₁ : ι₁ → Type v₁} [∀ i₁, Zero (β₁ i₁)] [∀ (i) (x : β₁ i), Decidable (x ≠ 0)] [∀ i, AddCommMonoid (β i)] {f : Π₀ i₁, β₁ i₁} {g : ∀ i₁, β₁ i₁ → Π₀ i, β i} {i₂ : ι} : (f.sum g) i₂ = f.sum fun i₁ b => g i₁ b i₂ := - (evalAddMonoidHom i₂ : (Π₀ i, β i) →+ β i₂).map_sum _ f.support + map_sum (evalAddMonoidHom i₂) _ f.support #align dfinsupp.sum_apply DFinsupp.sum_apply theorem support_sum {ι₁ : Type u₁} [DecidableEq ι₁] {β₁ : ι₁ → Type v₁} [∀ i₁, Zero (β₁ i₁)] @@ -1834,7 +1834,7 @@ theorem prod_mul [∀ i, AddCommMonoid (β i)] [∀ (i) (x : β i), Decidable (x @[to_additive (attr := simp)] theorem prod_inv [∀ i, AddCommMonoid (β i)] [∀ (i) (x : β i), Decidable (x ≠ 0)] [CommGroup γ] {f : Π₀ i, β i} {h : ∀ i, β i → γ} : (f.prod fun i b => (h i b)⁻¹) = (f.prod h)⁻¹ := - ((invMonoidHom : γ →* γ).map_prod _ f.support).symm + (map_prod (invMonoidHom : γ →* γ) _ f.support).symm #align dfinsupp.prod_inv DFinsupp.prod_inv #align dfinsupp.sum_neg DFinsupp.sum_neg @@ -2159,7 +2159,7 @@ theorem prod_subtypeDomain_index [∀ i, Zero (β i)] [∀ (i) (x : β i), Decid theorem subtypeDomain_sum [∀ i, AddCommMonoid (β i)] {s : Finset γ} {h : γ → Π₀ i, β i} {p : ι → Prop} [DecidablePred p] : (∑ c in s, h c).subtypeDomain p = ∑ c in s, (h c).subtypeDomain p := - (subtypeDomainAddMonoidHom β p).map_sum _ s + map_sum (subtypeDomainAddMonoidHom β p) _ s #align dfinsupp.subtype_domain_sum DFinsupp.subtypeDomain_sum theorem subtypeDomain_finsupp_sum {δ : γ → Type x} [DecidableEq γ] [∀ c, Zero (δ c)] diff --git a/Mathlib/Data/DFinsupp/Interval.lean b/Mathlib/Data/DFinsupp/Interval.lean index a08364767716e..e889b3ca8716c 100644 --- a/Mathlib/Data/DFinsupp/Interval.lean +++ b/Mathlib/Data/DFinsupp/Interval.lean @@ -198,7 +198,7 @@ variable [DecidableEq ι] [∀ i, DecidableEq (α i)] [∀ i, Lattice (α i)] [ [∀ i, LocallyFiniteOrder (α i)] (f g : Π₀ i, α i) theorem card_uIcc : (uIcc f g).card = ∏ i in f.support ∪ g.support, (uIcc (f i) (g i)).card := by - rw [←support_inf_union_support_sup]; exact card_Icc _ _ + rw [← support_inf_union_support_sup]; exact card_Icc _ _ #align dfinsupp.card_uIcc DFinsupp.card_uIcc end Lattice diff --git a/Mathlib/Data/DFinsupp/Order.lean b/Mathlib/Data/DFinsupp/Order.lean index 6f7dc2d19f6b8..1d91c7cd12abe 100644 --- a/Mathlib/Data/DFinsupp/Order.lean +++ b/Mathlib/Data/DFinsupp/Order.lean @@ -175,7 +175,7 @@ theorem le_iff : f ≤ g ↔ ∀ i ∈ f.support, f i ≤ g i := #align dfinsupp.le_iff DFinsupp.le_iff lemma support_monotone : Monotone (support (ι := ι) (β := α)) := - fun f g h a ha ↦ by rw [mem_support_iff, ←pos_iff_ne_zero] at ha ⊢; exact ha.trans_le (h _) + fun f g h a ha ↦ by rw [mem_support_iff, ← pos_iff_ne_zero] at ha ⊢; exact ha.trans_le (h _) lemma support_mono (hfg : f ≤ g) : f.support ⊆ g.support := support_monotone hfg diff --git a/Mathlib/Data/Fin/Basic.lean b/Mathlib/Data/Fin/Basic.lean index 404b0f960a7f6..957ee6a5f093e 100644 --- a/Mathlib/Data/Fin/Basic.lean +++ b/Mathlib/Data/Fin/Basic.lean @@ -511,8 +511,8 @@ instance nontrivial {n : ℕ} : Nontrivial (Fin (n + 2)) where theorem nontrivial_iff_two_le : Nontrivial (Fin n) ↔ 2 ≤ n := by rcases n with (_ | _ | n) <;> - simp [←Nat.one_eq_succ_zero, Fin.nontrivial, not_nontrivial, Nat.succ_le_iff] --- porting note: here and in the next lemma, had to use `←Nat.one_eq_succ_zero`. + simp [← Nat.one_eq_succ_zero, Fin.nontrivial, not_nontrivial, Nat.succ_le_iff] +-- porting note: here and in the next lemma, had to use `← Nat.one_eq_succ_zero`. #align fin.nontrivial_iff_two_le Fin.nontrivial_iff_two_le #align fin.subsingleton_iff_le_one Fin.subsingleton_iff_le_one diff --git a/Mathlib/Data/Fin/FlagRange.lean b/Mathlib/Data/Fin/FlagRange.lean index 343fc2fd9a12f..6d2f96ee9c456 100644 --- a/Mathlib/Data/Fin/FlagRange.lean +++ b/Mathlib/Data/Fin/FlagRange.lean @@ -33,7 +33,7 @@ theorem IsMaxChain.range_fin_of_covby (h0 : f 0 = ⊥) (hlast : f (.last n) = IsMaxChain (· ≤ ·) (range f) := by have hmono : Monotone f := Fin.monotone_iff_le_succ.2 fun k ↦ (hcovby k).1 refine ⟨hmono.isChain_range, fun t htc hbt ↦ hbt.antisymm fun x hx ↦ ?_⟩ - rw [mem_range]; by_contra' h + rw [mem_range]; by_contra! h suffices ∀ k, f k < x by simpa [hlast] using this (.last _) intro k induction k using Fin.induction with diff --git a/Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean b/Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean index 499a3fa7d53f3..d788eea08b447 100644 --- a/Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean +++ b/Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean @@ -91,7 +91,7 @@ theorem mem_antidiagonalTuple {n : ℕ} {k : ℕ} {x : Fin k → ℕ} : simp_rw [List.mem_bind, List.mem_map, List.Nat.mem_antidiagonal, Fin.cons_eq_cons, exists_eq_right_right, ih, @eq_comm _ _ (Prod.snd _), and_comm (a := Prod.snd _ = _), - ←Prod.mk.inj_iff (a₁ := Prod.fst _), exists_eq_right] + ← Prod.mk.inj_iff (a₁ := Prod.fst _), exists_eq_right] #align list.nat.mem_antidiagonal_tuple List.Nat.mem_antidiagonalTuple /-- The antidiagonal of `n` does not contain duplicate entries. -/ diff --git a/Mathlib/Data/Fin/Tuple/Sort.lean b/Mathlib/Data/Fin/Tuple/Sort.lean index 406a4c059e083..8ddb303d1017a 100644 --- a/Mathlib/Data/Fin/Tuple/Sort.lean +++ b/Mathlib/Data/Fin/Tuple/Sort.lean @@ -123,7 +123,7 @@ theorem lt_card_le_iff_apply_le_of_monotone [PartialOrder α] [DecidableRel (α j < Fintype.card {i // f i ≤ a} ↔ f j ≤ a := by suffices h1 : ∀ k : Fin m, (k < Fintype.card {i // f i ≤ a}) → f k ≤ a · refine ⟨h1 j, fun h ↦ ?_⟩ - by_contra' hc + by_contra! hc let p : Fin m → Prop := fun x ↦ f x ≤ a let q : Fin m → Prop := fun x ↦ x < Fintype.card {i // f i ≤ a} let q' : {i // f i ≤ a} → Prop := fun x ↦ q x diff --git a/Mathlib/Data/Finset/Basic.lean b/Mathlib/Data/Finset/Basic.lean index a27b38172e0ae..e8937430b36e8 100644 --- a/Mathlib/Data/Finset/Basic.lean +++ b/Mathlib/Data/Finset/Basic.lean @@ -1189,7 +1189,7 @@ theorem pair_comm (a b : α) : ({a, b} : Finset α) = {b, a} := -- Porting note: @[simp] can prove this theorem insert_idem (a : α) (s : Finset α) : insert a (insert a s) = insert a s := - ext fun x => by simp only [mem_insert, ←or_assoc, or_self_iff] + ext fun x => by simp only [mem_insert, ← or_assoc, or_self_iff] #align finset.insert_idem Finset.insert_idem @[simp] @@ -1226,7 +1226,7 @@ theorem insert_subset_insert (a : α) {s t : Finset α} (h : s ⊆ t) : insert a #align finset.insert_subset_insert Finset.insert_subset_insert @[simp] lemma insert_subset_insert_iff (ha : a ∉ s) : insert a s ⊆ insert a t ↔ s ⊆ t := by - simp_rw [←coe_subset]; simp [-coe_subset, ha] + simp_rw [← coe_subset]; simp [-coe_subset, ha] theorem insert_inj (ha : a ∉ s) : insert a s = insert b s ↔ a = b := ⟨fun h => eq_of_not_mem_of_mem_insert (h.subst <| mem_insert_self _ _) ha, congr_arg (insert · s)⟩ @@ -2121,8 +2121,8 @@ theorem inter_sdiff_self (s₁ s₂ : Finset α) : s₁ ∩ (s₂ \ s₁) = ∅ instance : GeneralizedBooleanAlgebra (Finset α) := { sup_inf_sdiff := fun x y => by - simp only [ext_iff, mem_union, mem_sdiff, inf_eq_inter, sup_eq_union, mem_inter, ←and_or_left, - em, and_true, implies_true] + simp only [ext_iff, mem_union, mem_sdiff, inf_eq_inter, sup_eq_union, mem_inter, + ← and_or_left, em, and_true, implies_true] inf_inf_sdiff := fun x y => by simp only [ext_iff, inter_sdiff_self, inter_empty, inter_assoc, false_iff_iff, inf_eq_inter, not_mem_empty, bot_eq_empty, not_false_iff, implies_true] } @@ -2908,7 +2908,7 @@ theorem filter_union (s₁ s₂ : Finset α) : (s₁ ∪ s₂).filter p = s₁.f #align finset.filter_union Finset.filter_union theorem filter_union_right (s : Finset α) : s.filter p ∪ s.filter q = s.filter fun x => p x ∨ q x := - ext fun x => by simp [mem_filter, mem_union, ←and_or_left] + ext fun x => by simp [mem_filter, mem_union, ← and_or_left] #align finset.filter_union_right Finset.filter_union_right theorem filter_mem_eq_inter {s t : Finset α} [∀ i, Decidable (i ∈ t)] : diff --git a/Mathlib/Data/Finset/Card.lean b/Mathlib/Data/Finset/Card.lean index 27b99d7d3fda3..64fddeb9862bb 100644 --- a/Mathlib/Data/Finset/Card.lean +++ b/Mathlib/Data/Finset/Card.lean @@ -358,7 +358,7 @@ theorem card_le_card_of_inj_on {t : Finset β} (f : α → β) (hf : ∀ a ∈ s theorem exists_ne_map_eq_of_card_lt_of_maps_to {t : Finset β} (hc : t.card < s.card) {f : α → β} (hf : ∀ a ∈ s, f a ∈ t) : ∃ x ∈ s, ∃ y ∈ s, x ≠ y ∧ f x = f y := by classical - by_contra' hz + by_contra! hz refine' hc.not_le (card_le_card_of_inj_on f hf _) intro x hx y hy contrapose @@ -473,7 +473,7 @@ theorem card_sdiff_add_card : (s \ t).card + t.card = (s ∪ t).card := by #align finset.card_sdiff_add_card Finset.card_sdiff_add_card lemma card_sdiff_comm (h : s.card = t.card) : (s \ t).card = (t \ s).card := - add_left_injective t.card $ by simp_rw [card_sdiff_add_card, ←h, card_sdiff_add_card, union_comm] + add_left_injective t.card $ by simp_rw [card_sdiff_add_card, ← h, card_sdiff_add_card, union_comm] @[simp] lemma card_sdiff_add_card_inter (s t : Finset α) : @@ -507,7 +507,7 @@ theorem exists_intermediate_set {A B : Finset α} (i : ℕ) (h₁ : i + card B apply Nat.succ_pos have z : i + card B + k = card (erase A a) := by rw [card_erase_of_mem (mem_sdiff.1 ha).1, ← h, - Nat.add_sub_assoc (Nat.one_le_iff_ne_zero.mpr k.succ_ne_zero), ←pred_eq_sub_one, + Nat.add_sub_assoc (Nat.one_le_iff_ne_zero.mpr k.succ_ne_zero), ← pred_eq_sub_one, k.pred_succ] have : B ⊆ A.erase a := by rintro t th @@ -685,7 +685,7 @@ lemma covby_iff_card_sdiff_eq_one : t ⋖ s ↔ t ⊆ s ∧ (s \ t).card = 1 := · simp_rw [card_eq_one] rintro ⟨hts, a, ha⟩ refine ⟨a, (mem_sdiff.1 $ superset_of_eq ha $ mem_singleton_self _).2, ?_⟩ - rw [insert_eq, ←ha, sdiff_union_of_subset hts] + rw [insert_eq, ← ha, sdiff_union_of_subset hts] end DecidableEq diff --git a/Mathlib/Data/Finset/Functor.lean b/Mathlib/Data/Finset/Functor.lean index 9809b802743a6..8cd398b6f147b 100644 --- a/Mathlib/Data/Finset/Functor.lean +++ b/Mathlib/Data/Finset/Functor.lean @@ -165,7 +165,7 @@ instance : LawfulMonad Finset := bind_pure_comp := fun f s => sup_singleton'' _ _ bind_map := fun t s => rfl pure_bind := fun t s => sup_singleton - bind_assoc := fun s f g => by simp only [bind, ←sup_biUnion, sup_eq_biUnion, biUnion_biUnion] } + bind_assoc := fun s f g => by simp only [bind, ← sup_biUnion, sup_eq_biUnion, biUnion_biUnion] } end Monad diff --git a/Mathlib/Data/Finset/Image.lean b/Mathlib/Data/Finset/Image.lean index f24b7f0010bb6..897a27f4cf21c 100644 --- a/Mathlib/Data/Finset/Image.lean +++ b/Mathlib/Data/Finset/Image.lean @@ -490,7 +490,7 @@ theorem image_subset_image_iff {t : Finset α} (hf : Injective f) : #align finset.image_subset_image_iff Finset.image_subset_image_iff lemma image_ssubset_image {t : Finset α} (hf : Injective f) : s.image f ⊂ t.image f ↔ s ⊂ t := by - simp_rw [←lt_iff_ssubset] + simp_rw [← lt_iff_ssubset] exact lt_iff_lt_of_le_iff_le' (image_subset_image_iff hf) (image_subset_image_iff hf) theorem coe_image_subset_range : ↑(s.image f) ⊆ Set.range f := @@ -716,7 +716,7 @@ theorem filterMap_some : s.filterMap some (by simp) = s := theorem filterMap_mono (h : s ⊆ t) : filterMap f s f_inj ⊆ filterMap f t f_inj := by - rw [←val_le_iff] at h ⊢ + rw [← val_le_iff] at h ⊢ exact Multiset.filterMap_le_filterMap f h end FilterMap diff --git a/Mathlib/Data/Finset/Lattice.lean b/Mathlib/Data/Finset/Lattice.lean index 59f8e62a454f5..f65271004fe52 100644 --- a/Mathlib/Data/Finset/Lattice.lean +++ b/Mathlib/Data/Finset/Lattice.lean @@ -901,7 +901,7 @@ theorem _root_.map_finset_sup' [SemilatticeSup β] [SupHomClass F α β] (f : F) theorem sup'_image [DecidableEq β] {s : Finset γ} {f : γ → β} (hs : (s.image f).Nonempty) (g : β → α) (hs' : s.Nonempty := (Nonempty.image_iff _).1 hs) : (s.image f).sup' hs g = s.sup' hs' (g ∘ f) := by - rw [←WithBot.coe_eq_coe]; simp only [coe_sup', sup_image, WithBot.coe_sup]; rfl + rw [← WithBot.coe_eq_coe]; simp only [coe_sup', sup_image, WithBot.coe_sup]; rfl #align finset.sup'_image Finset.sup'_image @[simp] diff --git a/Mathlib/Data/Finset/NoncommProd.lean b/Mathlib/Data/Finset/NoncommProd.lean index 4d6d7a48fe920..be569f1cdd022 100644 --- a/Mathlib/Data/Finset/NoncommProd.lean +++ b/Mathlib/Data/Finset/NoncommProd.lean @@ -217,7 +217,7 @@ theorem noncommProd_eq_prod {α : Type*} [CommMonoid α] (s : Multiset α) : #align multiset.noncomm_prod_eq_prod Multiset.noncommProd_eq_prod #align multiset.noncomm_sum_eq_sum Multiset.noncommSum_eq_sum -@[to_additive noncommSum_addCommute] +@[to_additive] theorem noncommProd_commute (s : Multiset α) (comm) (y : α) (h : ∀ x ∈ s, Commute y x) : Commute y (s.noncommProd comm) := by induction s using Quotient.inductionOn @@ -366,7 +366,7 @@ theorem noncommProd_eq_pow_card (s : Finset α) (f : α → β) (comm) (m : β) #align finset.noncomm_prod_eq_pow_card Finset.noncommProd_eq_pow_card #align finset.noncomm_sum_eq_card_nsmul Finset.noncommSum_eq_card_nsmul -@[to_additive noncommSum_addCommute] +@[to_additive] theorem noncommProd_commute (s : Finset α) (f : α → β) (comm) (y : β) (h : ∀ x ∈ s, Commute y (f x)) : Commute y (s.noncommProd f comm) := by apply Multiset.noncommProd_commute diff --git a/Mathlib/Data/Finset/PImage.lean b/Mathlib/Data/Finset/PImage.lean index 37a86f78947b3..b4044caf67596 100644 --- a/Mathlib/Data/Finset/PImage.lean +++ b/Mathlib/Data/Finset/PImage.lean @@ -82,8 +82,8 @@ theorem pimage_some (s : Finset α) (f : α → β) [∀ x, Decidable (Part.some theorem pimage_congr (h₁ : s = t) (h₂ : ∀ x ∈ t, f x = g x) : s.pimage f = t.pimage g := by subst s ext y - -- Porting note: `←exists_prop` required because `∃ x ∈ s, p x` is defined differently - simp (config := { contextual := true }) only [mem_pimage, ←exists_prop, h₂] + -- Porting note: `← exists_prop` required because `∃ x ∈ s, p x` is defined differently + simp (config := { contextual := true }) only [mem_pimage, ← exists_prop, h₂] #align finset.pimage_congr Finset.pimage_congr /-- Rewrite `s.pimage f` in terms of `Finset.filter`, `Finset.attach`, and `Finset.image`. -/ @@ -93,8 +93,8 @@ theorem pimage_eq_image_filter : s.pimage f = (f x).get (mem_filter.mp x.coe_prop).2 := by ext x simp [Part.mem_eq, And.exists] - -- Porting note: `←exists_prop` required because `∃ x ∈ s, p x` is defined differently - simp only [←exists_prop] + -- Porting note: `← exists_prop` required because `∃ x ∈ s, p x` is defined differently + simp only [← exists_prop] #align finset.pimage_eq_image_filter Finset.pimage_eq_image_filter theorem pimage_union [DecidableEq α] : (s ∪ t).pimage f = s.pimage f ∪ t.pimage f := diff --git a/Mathlib/Data/Finset/Sups.lean b/Mathlib/Data/Finset/Sups.lean index c4704148c0923..585619caf3107 100644 --- a/Mathlib/Data/Finset/Sups.lean +++ b/Mathlib/Data/Finset/Sups.lean @@ -187,13 +187,13 @@ lemma map_sups (f : F) (hf) (s t : Finset α) : lemma subset_sups_self : s ⊆ s ⊻ s := fun _a ha ↦ mem_sups.2 ⟨_, ha, _, ha, sup_idem⟩ lemma sups_subset_self : s ⊻ s ⊆ s ↔ SupClosed (s : Set α) := sups_subset_iff -@[simp] lemma sups_eq_self : s ⊻ s = s ↔ SupClosed (s : Set α) := by simp [←coe_inj] +@[simp] lemma sups_eq_self : s ⊻ s = s ↔ SupClosed (s : Set α) := by simp [← coe_inj] @[simp] lemma univ_sups_univ [Fintype α] : (univ : Finset α) ⊻ univ = univ := by simp lemma filter_sups_le [@DecidableRel α (· ≤ ·)] (s t : Finset α) (a : α) : (s ⊻ t).filter (· ≤ a) = s.filter (· ≤ a) ⊻ t.filter (· ≤ a) := by - simp only [←coe_inj, coe_filter, coe_sups, ←mem_coe, Set.sep_sups_le] + simp only [← coe_inj, coe_filter, coe_sups, ← mem_coe, Set.sep_sups_le] variable (s t u) @@ -371,13 +371,13 @@ lemma map_infs (f : F) (hf) (s t : Finset α) : lemma subset_infs_self : s ⊆ s ⊼ s := fun _a ha ↦ mem_infs.2 ⟨_, ha, _, ha, inf_idem⟩ lemma infs_self_subset : s ⊼ s ⊆ s ↔ InfClosed (s : Set α) := infs_subset_iff -@[simp] lemma infs_self : s ⊼ s = s ↔ InfClosed (s : Set α) := by simp [←coe_inj] +@[simp] lemma infs_self : s ⊼ s = s ↔ InfClosed (s : Set α) := by simp [← coe_inj] @[simp] lemma univ_infs_univ [Fintype α] : (univ : Finset α) ⊼ univ = univ := by simp lemma filter_infs_le [@DecidableRel α (· ≤ ·)] (s t : Finset α) (a : α) : (s ⊼ t).filter (a ≤ ·) = s.filter (a ≤ ·) ⊼ t.filter (a ≤ ·) := by - simp only [←coe_inj, coe_filter, coe_infs, ←mem_coe, Set.sep_infs_le] + simp only [← coe_inj, coe_filter, coe_infs, ← mem_coe, Set.sep_infs_le] variable (s t u) @@ -447,7 +447,7 @@ variable {𝒜 ℬ : Finset (Finset α)} {s t : Finset α} {a : α} ext u simp only [mem_sups, mem_powerset, le_eq_subset, sup_eq_union] refine ⟨fun h ↦ ⟨_, inter_subset_left _ u, _, inter_subset_left _ u, ?_⟩, ?_⟩ - · rwa [←inter_distrib_right, inter_eq_right] + · rwa [← inter_distrib_right, inter_eq_right] · rintro ⟨v, hv, w, hw, rfl⟩ exact union_subset_union hv hw @@ -455,15 +455,15 @@ variable {𝒜 ℬ : Finset (Finset α)} {s t : Finset α} {a : α} ext u simp only [mem_infs, mem_powerset, le_eq_subset, inf_eq_inter] refine ⟨fun h ↦ ⟨_, inter_subset_left _ u, _, inter_subset_left _ u, ?_⟩, ?_⟩ - · rwa [←inter_inter_distrib_right, inter_eq_right] + · rwa [← inter_inter_distrib_right, inter_eq_right] · rintro ⟨v, hv, w, hw, rfl⟩ exact inter_subset_inter hv hw @[simp] lemma powerset_sups_powerset_self (s : Finset α) : - s.powerset ⊻ s.powerset = s.powerset := by simp [←powerset_union] + s.powerset ⊻ s.powerset = s.powerset := by simp [← powerset_union] @[simp] lemma powerset_infs_powerset_self (s : Finset α) : - s.powerset ⊼ s.powerset = s.powerset := by simp [←powerset_inter] + s.powerset ⊼ s.powerset = s.powerset := by simp [← powerset_inter] lemma union_mem_sups : s ∈ 𝒜 → t ∈ ℬ → s ∪ t ∈ 𝒜 ⊻ ℬ := sup_mem_sups lemma inter_mem_infs : s ∈ 𝒜 → t ∈ ℬ → s ∩ t ∈ 𝒜 ⊼ ℬ := inf_mem_infs @@ -706,7 +706,7 @@ open FinsetFamily variable {s t} {a b c : α} @[simp] lemma mem_compls : a ∈ sᶜˢ ↔ aᶜ ∈ s := by - rw [Iff.comm, ←mem_map' ⟨compl, compl_injective⟩, Embedding.coeFn_mk, compl_compl, compls] + rw [Iff.comm, ← mem_map' ⟨compl, compl_injective⟩, Embedding.coeFn_mk, compl_compl, compls] variable (s t) @@ -725,7 +725,7 @@ lemma exists_compls_iff {p : α → Prop} : (∃ a ∈ sᶜˢ, p a) ↔ ∃ a @[simp] lemma compls_compls (s : Finset α) : sᶜˢᶜˢ = s := by ext; simp -lemma compls_subset_iff : sᶜˢ ⊆ t ↔ s ⊆ tᶜˢ := by rw [←compls_subset_compls, compls_compls] +lemma compls_subset_iff : sᶜˢ ⊆ t ↔ s ⊆ tᶜˢ := by rw [← compls_subset_compls, compls_compls] @[simp] lemma compls_nonempty : sᶜˢ.Nonempty ↔ s.Nonempty := map_nonempty @@ -739,10 +739,10 @@ protected alias ⟨Nonempty.of_compls, Nonempty.compls⟩ := compls_nonempty @[simp] lemma compls_inter (s t : Finset α) : (s ∩ t)ᶜˢ = sᶜˢ ∩ tᶜˢ := map_inter _ _ @[simp] lemma compls_infs (s t : Finset α) : (s ⊼ t)ᶜˢ = sᶜˢ ⊻ tᶜˢ := by - simp_rw [←image_compl]; exact image_image₂_distrib λ _ _ ↦ compl_inf + simp_rw [← image_compl]; exact image_image₂_distrib λ _ _ ↦ compl_inf @[simp] lemma compls_sups (s t : Finset α) : (s ⊻ t)ᶜˢ = sᶜˢ ⊼ tᶜˢ := by - simp_rw [←image_compl]; exact image_image₂_distrib λ _ _ ↦ compl_sup + simp_rw [← image_compl]; exact image_image₂_distrib λ _ _ ↦ compl_sup @[simp] lemma infs_compls_eq_diffs (s t : Finset α) : s ⊼ tᶜˢ = s \\ t := by ext; simp [sdiff_eq]; aesop @@ -751,7 +751,7 @@ protected alias ⟨Nonempty.of_compls, Nonempty.compls⟩ := compls_nonempty rw [infs_comm, infs_compls_eq_diffs] @[simp] lemma diffs_compls_eq_infs (s t : Finset α) : s \\ tᶜˢ = s ⊼ t := by - rw [←infs_compls_eq_diffs, compls_compls] + rw [← infs_compls_eq_diffs, compls_compls] variable [Fintype α] {𝒜 : Finset (Finset α)} {n : ℕ} diff --git a/Mathlib/Data/Finsupp/AList.lean b/Mathlib/Data/Finsupp/AList.lean index 98ed6de508117..0e702c1b74605 100644 --- a/Mathlib/Data/Finsupp/AList.lean +++ b/Mathlib/Data/Finsupp/AList.lean @@ -117,7 +117,7 @@ theorem insert_lookupFinsupp [DecidableEq α] (l : AList fun _x : α => M) (a : theorem singleton_lookupFinsupp (a : α) (m : M) : (singleton a m).lookupFinsupp = Finsupp.single a m := by classical - -- porting note: was `simp [←AList.insert_empty]` but timeout issues + -- porting note: was `simp [← AList.insert_empty]` but timeout issues simp only [← AList.insert_empty, insert_lookupFinsupp, empty_lookupFinsupp, Finsupp.zero_update] #align alist.singleton_lookup_finsupp AList.singleton_lookupFinsupp diff --git a/Mathlib/Data/Finsupp/Antidiagonal.lean b/Mathlib/Data/Finsupp/Antidiagonal.lean index c3a9c042631c3..249a21f269ca9 100644 --- a/Mathlib/Data/Finsupp/Antidiagonal.lean +++ b/Mathlib/Data/Finsupp/Antidiagonal.lean @@ -72,7 +72,7 @@ theorem antidiagonal_single (a : α) (n : ℕ) : constructor · intro h refine ⟨x a, y a, FunLike.congr_fun h a |>.trans single_eq_same, ?_⟩ - simp_rw [FunLike.ext_iff, ←forall_and] + simp_rw [FunLike.ext_iff, ← forall_and] intro i replace h := FunLike.congr_fun h i simp_rw [single_apply, Finsupp.add_apply] at h ⊢ diff --git a/Mathlib/Data/Finsupp/Basic.lean b/Mathlib/Data/Finsupp/Basic.lean index 65b18080ddd11..272733f82aa35 100644 --- a/Mathlib/Data/Finsupp/Basic.lean +++ b/Mathlib/Data/Finsupp/Basic.lean @@ -241,7 +241,7 @@ theorem mapRange_multiset_sum (f : F) (m : Multiset (α →₀ M)) : theorem mapRange_finset_sum (f : F) (s : Finset ι) (g : ι → α →₀ M) : mapRange f (map_zero f) (∑ x in s, g x) = ∑ x in s, mapRange f (map_zero f) (g x) := - (mapRange.addMonoidHom (f : M →+ N) : (α →₀ _) →+ _).map_sum _ _ + map_sum (mapRange.addMonoidHom (f : M →+ N)) _ _ #align finsupp.map_range_finset_sum Finsupp.mapRange_finset_sum /-- `Finsupp.mapRange.AddMonoidHom` as an equiv. -/ @@ -528,7 +528,7 @@ theorem mapDomain.addMonoidHom_comp (f : β → γ) (g : α → β) : theorem mapDomain_finset_sum {f : α → β} {s : Finset ι} {v : ι → α →₀ M} : mapDomain f (∑ i in s, v i) = ∑ i in s, mapDomain f (v i) := - (mapDomain.addMonoidHom f : (α →₀ M) →+ β →₀ M).map_sum _ _ + map_sum (mapDomain.addMonoidHom f) _ _ #align finsupp.map_domain_finset_sum Finsupp.mapDomain_finset_sum theorem mapDomain_sum [Zero N] {f : α → β} {s : α →₀ N} {v : α → N → α →₀ M} : @@ -1099,7 +1099,7 @@ variable [AddCommMonoid M] {p : α → Prop} theorem subtypeDomain_sum {s : Finset ι} {h : ι → α →₀ M} : (∑ c in s, h c).subtypeDomain p = ∑ c in s, (h c).subtypeDomain p := - (subtypeDomainAddMonoidHom : _ →+ Subtype p →₀ M).map_sum _ s + map_sum subtypeDomainAddMonoidHom _ s #align finsupp.subtype_domain_sum Finsupp.subtypeDomain_sum theorem subtypeDomain_finsupp_sum [Zero N] {s : β →₀ N} {h : β → N → α →₀ M} : @@ -1109,7 +1109,7 @@ theorem subtypeDomain_finsupp_sum [Zero N] {s : β →₀ N} {h : β → N → theorem filter_sum (s : Finset ι) (f : ι → α →₀ M) : (∑ a in s, f a).filter p = ∑ a in s, filter p (f a) := - (filterAddHom p : (α →₀ M) →+ _).map_sum f s + map_sum (filterAddHom p) f s #align finsupp.filter_sum Finsupp.filter_sum theorem filter_eq_sum (p : α → Prop) [D : DecidablePred p] (f : α →₀ M) : @@ -1706,7 +1706,7 @@ def restrictSupportEquiv (s : Set α) (M : Type*) [AddCommMonoid M] : by_cases h : a ∈ s · lift a to s using h exact embDomain_apply _ _ _ - rw [embDomain_notin_range, eq_comm, ←Finsupp.not_mem_support_iff] + rw [embDomain_notin_range, eq_comm, ← Finsupp.not_mem_support_iff] · exact fun hs => h <| hf hs · simp [h] right_inv f := ext <| embDomain_apply _ f diff --git a/Mathlib/Data/Finsupp/Order.lean b/Mathlib/Data/Finsupp/Order.lean index d0deb270f8712..89fe37464be1d 100644 --- a/Mathlib/Data/Finsupp/Order.lean +++ b/Mathlib/Data/Finsupp/Order.lean @@ -174,7 +174,7 @@ theorem le_iff (f g : ι →₀ α) : f ≤ g ↔ ∀ i ∈ f.support, f i ≤ g #align finsupp.le_iff Finsupp.le_iff lemma support_monotone : Monotone (support (α := ι) (M := α)) := - fun f g h a ha ↦ by rw [mem_support_iff, ←pos_iff_ne_zero] at ha ⊢; exact ha.trans_le (h _) + fun f g h a ha ↦ by rw [mem_support_iff, ← pos_iff_ne_zero] at ha ⊢; exact ha.trans_le (h _) lemma support_mono (hfg : f ≤ g) : f.support ⊆ g.support := support_monotone hfg diff --git a/Mathlib/Data/Fintype/Basic.lean b/Mathlib/Data/Fintype/Basic.lean index bee30a7edd088..9c075e7a1c66c 100644 --- a/Mathlib/Data/Fintype/Basic.lean +++ b/Mathlib/Data/Fintype/Basic.lean @@ -1227,7 +1227,7 @@ theorem count_univ [DecidableEq α] (a : α) : count a Finset.univ.val = 1 := @[simp] theorem map_univ_val_equiv (e : α ≃ β) : map e univ.val = univ.val := by - rw [←congr_arg Finset.val (Finset.map_univ_equiv e), Finset.map_val, Equiv.coe_toEmbedding] + rw [← congr_arg Finset.val (Finset.map_univ_equiv e), Finset.map_val, Equiv.coe_toEmbedding] /-- For functions on finite sets, they are bijections iff they map universes into universes. -/ @[simp] diff --git a/Mathlib/Data/Fintype/Card.lean b/Mathlib/Data/Fintype/Card.lean index 1ef85ff6eaf8f..dd578060bf36f 100644 --- a/Mathlib/Data/Fintype/Card.lean +++ b/Mathlib/Data/Fintype/Card.lean @@ -1199,7 +1199,7 @@ See also: `Finite.exists_ne_map_eq_of_infinite` theorem Finite.exists_infinite_fiber [Infinite α] [Finite β] (f : α → β) : ∃ y : β, Infinite (f ⁻¹' {y}) := by classical - by_contra' hf + by_contra! hf cases nonempty_fintype β haveI := fun y => fintypeOfNotInfinite <| hf y let key : Fintype α := @@ -1313,7 +1313,7 @@ private theorem card_univ_pos (α : Type*) [Fintype α] [Nonempty α] : -- unsafe def positivity_finset_card : expr → tactic strictness -- | q(Finset.card $(s)) => do -- let p --- ←-- TODO: Partial decision procedure for `Finset.nonempty` +-- ← -- TODO: Partial decision procedure for `Finset.nonempty` -- to_expr -- ``(Finset.Nonempty $(s)) >>= -- find_assumption diff --git a/Mathlib/Data/Fintype/CardEmbedding.lean b/Mathlib/Data/Fintype/CardEmbedding.lean index c6a9731d7fba5..a2fb358165c70 100644 --- a/Mathlib/Data/Fintype/CardEmbedding.lean +++ b/Mathlib/Data/Fintype/CardEmbedding.lean @@ -46,7 +46,7 @@ theorem card_embedding_eq {α β : Type*} [Fintype α] [Fintype β] [emb : Finty · classical dsimp only at ih rw [card_option, Nat.descFactorial_succ, card_congr (Embedding.optionEmbeddingEquiv γ β), - card_sigma, ←ih] + card_sigma, ← ih] simp only [Fintype.card_compl_set, Fintype.card_range, Finset.sum_const, Finset.card_univ, smul_eq_mul, mul_comm] #align fintype.card_embedding_eq Fintype.card_embedding_eq diff --git a/Mathlib/Data/Holor.lean b/Mathlib/Data/Holor.lean index eed50ec61d359..eda0330598e3d 100644 --- a/Mathlib/Data/Holor.lean +++ b/Mathlib/Data/Holor.lean @@ -196,7 +196,8 @@ theorem mul_assoc0 [Semigroup α] (x : Holor α ds₁) (y : Holor α ds₂) (z : funext fun t : HolorIndex (ds₁ ++ ds₂ ++ ds₃) => by rw [assocLeft] unfold mul - rw [mul_assoc, ←HolorIndex.take_take, ←HolorIndex.drop_take, ←HolorIndex.drop_drop, cast_type] + rw [mul_assoc, ← HolorIndex.take_take, ← HolorIndex.drop_take, ← HolorIndex.drop_drop, + cast_type] rfl rw [append_assoc] #align holor.mul_assoc0 Holor.mul_assoc0 diff --git a/Mathlib/Data/Int/Bitwise.lean b/Mathlib/Data/Int/Bitwise.lean index f776c5598e108..9ddfb504c328c 100644 --- a/Mathlib/Data/Int/Bitwise.lean +++ b/Mathlib/Data/Int/Bitwise.lean @@ -61,7 +61,7 @@ theorem bodd_negOfNat (n : ℕ) : bodd (negOfNat n) = n.bodd := by theorem bodd_neg (n : ℤ) : bodd (-n) = bodd n := by cases n with | ofNat => - rw [←negOfNat_eq, bodd_negOfNat] + rw [← negOfNat_eq, bodd_negOfNat] simp | negSucc n => rw [neg_negSucc, bodd_coe, Nat.bodd_succ] @@ -77,8 +77,8 @@ theorem bodd_add (m n : ℤ) : bodd (m + n) = xor (bodd m) (bodd n) := by cases' n with n n <;> simp only [ofNat_eq_coe, ofNat_add_negSucc, negSucc_add_ofNat, negSucc_add_negSucc, bodd_subNatNat] <;> - simp only [negSucc_coe, bodd_neg, bodd_coe, ←Nat.bodd_add, Bool.xor_comm, ←Nat.cast_add] - rw [←Nat.succ_add, add_assoc] + simp only [negSucc_coe, bodd_neg, bodd_coe, ← Nat.bodd_add, Bool.xor_comm, ← Nat.cast_add] + rw [← Nat.succ_add, add_assoc] -- Porting note: Heavily refactored proof, used to work all with `simp`: -- `by cases m with m m; cases n with n n; unfold has_add.add;` -- `simp [int.add, -of_nat_eq_coe, bool.xor_comm]` @@ -89,7 +89,7 @@ theorem bodd_mul (m n : ℤ) : bodd (m * n) = (bodd m && bodd n) := by cases' m with m m <;> cases' n with n n <;> simp only [ofNat_eq_coe, ofNat_mul_negSucc, negSucc_mul_ofNat, ofNat_mul_ofNat, negSucc_mul_negSucc] <;> - simp only [negSucc_coe, bodd_neg, bodd_coe, ←Nat.bodd_mul] + simp only [negSucc_coe, bodd_neg, bodd_coe, ← Nat.bodd_mul] -- Porting note: Heavily refactored proof, used to be: -- `by cases m with m m; cases n with n n;` -- `simp [← int.mul_def, int.mul, -of_nat_eq_coe, bool.xor_comm]` @@ -407,15 +407,15 @@ attribute [local simp] Int.zero_div theorem shiftLeft_add : ∀ (m : ℤ) (n : ℕ) (k : ℤ), m <<< (n + k) = (m <<< (n : ℤ)) <<< k | (m : ℕ), n, (k : ℕ) => - congr_arg ofNat (by simp [Nat.pow_add, mul_assoc]) + congr_arg ofNat (by simp [Nat.shiftLeft_eq, Nat.pow_add, mul_assoc]) | -[m+1], n, (k : ℕ) => congr_arg negSucc (Nat.shiftLeft'_add _ _ _ _) | (m : ℕ), n, -[k+1] => subNatNat_elim n k.succ (fun n k i => (↑m) <<< i = (Nat.shiftLeft' false m n) >>> k) (fun (i n : ℕ) => - by dsimp; simp [- Nat.shiftLeft_eq, ← Nat.shiftLeft_sub _ , add_tsub_cancel_left]) + by dsimp; simp [← Nat.shiftLeft_sub _ , add_tsub_cancel_left]) fun i n => by dsimp - simp [- Nat.shiftLeft_eq, Nat.shiftLeft_zero, Nat.shiftRight_add, ← Nat.shiftLeft_sub] + simp [Nat.shiftLeft_zero, Nat.shiftRight_add, ← Nat.shiftLeft_sub] rfl | -[m+1], n, -[k+1] => subNatNat_elim n k.succ @@ -433,7 +433,7 @@ theorem shiftLeft_sub (m : ℤ) (n : ℕ) (k : ℤ) : m <<< (n - k) = (m <<< (n #align int.shiftl_sub Int.shiftLeft_sub theorem shiftLeft_eq_mul_pow : ∀ (m : ℤ) (n : ℕ), m <<< (n : ℤ) = m * (2 ^ n : ℕ) - | (m : ℕ), _ => congr_arg ((↑) : ℕ → ℤ) (by simp) + | (m : ℕ), _ => congr_arg ((↑) : ℕ → ℤ) (by simp [Nat.shiftLeft_eq]) | -[_+1], _ => @congr_arg ℕ ℤ _ _ (fun i => -i) (Nat.shiftLeft'_tt_eq_mul_pow _ _) #align int.shiftl_eq_mul_pow Int.shiftLeft_eq_mul_pow @@ -445,7 +445,7 @@ theorem shiftRight_eq_div_pow : ∀ (m : ℤ) (n : ℕ), m >>> (n : ℤ) = m / ( #align int.shiftr_eq_div_pow Int.shiftRight_eq_div_pow theorem one_shiftLeft (n : ℕ) : 1 <<< (n : ℤ) = (2 ^ n : ℕ) := - congr_arg ((↑) : ℕ → ℤ) (by simp) + congr_arg ((↑) : ℕ → ℤ) (by simp [Nat.shiftLeft_eq]) #align int.one_shiftl Int.one_shiftLeft @[simp] diff --git a/Mathlib/Data/Int/Cast/Lemmas.lean b/Mathlib/Data/Int/Cast/Lemmas.lean index 12c59b7d5b16a..557ffd7adac5e 100644 --- a/Mathlib/Data/Int/Cast/Lemmas.lean +++ b/Mathlib/Data/Int/Cast/Lemmas.lean @@ -47,7 +47,7 @@ theorem coe_nat_succ_pos (n : ℕ) : 0 < (n.succ : ℤ) := #align int.coe_nat_succ_pos Int.coe_nat_succ_pos lemma toNat_lt' {a : ℤ} {b : ℕ} (hb : b ≠ 0) : a.toNat < b ↔ a < b := by - rw [←toNat_lt_toNat, toNat_coe_nat]; exact coe_nat_pos.2 hb.bot_lt + rw [← toNat_lt_toNat, toNat_coe_nat]; exact coe_nat_pos.2 hb.bot_lt #align int.to_nat_lt Int.toNat_lt' lemma natMod_lt {a : ℤ} {b : ℕ} (hb : b ≠ 0) : a.natMod b < b := diff --git a/Mathlib/Data/Int/Dvd/Basic.lean b/Mathlib/Data/Int/Dvd/Basic.lean index 0a8b5d3191992..9f905edfecbc9 100644 --- a/Mathlib/Data/Int/Dvd/Basic.lean +++ b/Mathlib/Data/Int/Dvd/Basic.lean @@ -31,11 +31,11 @@ theorem coe_nat_dvd {m n : ℕ} : (↑m : ℤ) ∣ ↑n ↔ m ∣ n := #align int.coe_nat_dvd Int.coe_nat_dvd theorem coe_nat_dvd_left {n : ℕ} {z : ℤ} : (↑n : ℤ) ∣ z ↔ n ∣ z.natAbs := by - rcases natAbs_eq z with (eq | eq) <;> rw [eq] <;> simp [←coe_nat_dvd] + rcases natAbs_eq z with (eq | eq) <;> rw [eq] <;> simp [← coe_nat_dvd] #align int.coe_nat_dvd_left Int.coe_nat_dvd_left theorem coe_nat_dvd_right {n : ℕ} {z : ℤ} : z ∣ (↑n : ℤ) ↔ z.natAbs ∣ n := by - rcases natAbs_eq z with (eq | eq) <;> rw [eq] <;> simp [←coe_nat_dvd] + rcases natAbs_eq z with (eq | eq) <;> rw [eq] <;> simp [← coe_nat_dvd] #align int.coe_nat_dvd_right Int.coe_nat_dvd_right #align int.le_of_dvd Int.le_of_dvd diff --git a/Mathlib/Data/Int/Order/Basic.lean b/Mathlib/Data/Int/Order/Basic.lean index 9e6ca9ec5346c..16546f32a5770 100644 --- a/Mathlib/Data/Int/Order/Basic.lean +++ b/Mathlib/Data/Int/Order/Basic.lean @@ -62,7 +62,7 @@ theorem abs_eq_natAbs : ∀ a : ℤ, |a| = natAbs a #align int.coe_nat_abs Int.coe_natAbs lemma _root_.Nat.cast_natAbs {α : Type*} [AddGroupWithOne α] (n : ℤ) : (n.natAbs : α) = |n| := - by rw [←coe_natAbs, Int.cast_ofNat] + by rw [← coe_natAbs, Int.cast_ofNat] #align nat.cast_nat_abs Nat.cast_natAbs theorem natAbs_abs (a : ℤ) : natAbs |a| = natAbs a := by rw [abs_eq_natAbs]; rfl @@ -120,7 +120,7 @@ theorem le_sub_one_iff {a b : ℤ} : a ≤ b - 1 ↔ a < b := theorem abs_lt_one_iff {a : ℤ} : |a| < 1 ↔ a = 0 := ⟨fun a0 => by let ⟨hn, hp⟩ := abs_lt.mp a0 - rw [←zero_add 1, lt_add_one_iff] at hp + rw [← zero_add 1, lt_add_one_iff] at hp -- Defeq abuse: `hn : -1 < a` but should be `hn : 0 λ a`. exact hp.antisymm hn, fun a0 => (abs_eq_zero.mpr a0).le.trans_lt zero_lt_one⟩ diff --git a/Mathlib/Data/Int/Sqrt.lean b/Mathlib/Data/Int/Sqrt.lean index 0178909058a11..a8667c8ddf007 100644 --- a/Mathlib/Data/Int/Sqrt.lean +++ b/Mathlib/Data/Int/Sqrt.lean @@ -40,6 +40,6 @@ theorem sqrt_nonneg (n : ℤ) : 0 ≤ sqrt n := /-- `IsSquare` can be decided on `ℤ` by checking against the square root. -/ instance : DecidablePred (IsSquare : ℤ → Prop) := fun m => decidable_of_iff' (sqrt m * sqrt m = m) <| by - simp_rw [←exists_mul_self m, IsSquare, eq_comm] + simp_rw [← exists_mul_self m, IsSquare, eq_comm] end Int diff --git a/Mathlib/Data/List/Basic.lean b/Mathlib/Data/List/Basic.lean index 4835d2d21dc3a..13e247bc943e9 100644 --- a/Mathlib/Data/List/Basic.lean +++ b/Mathlib/Data/List/Basic.lean @@ -2540,7 +2540,7 @@ variable {f : β → α → β} {b : β} {a : α} {l : List α} theorem length_scanl : ∀ a l, length (scanl f a l) = l.length + 1 | a, [] => rfl | a, x :: l => by - rw [scanl, length_cons, length_cons, ←succ_eq_add_one, congr_arg succ] + rw [scanl, length_cons, length_cons, ← succ_eq_add_one, congr_arg succ] exact length_scanl _ _ #align list.length_scanl List.length_scanl @@ -2752,7 +2752,7 @@ theorem foldlM_eq_foldl (f : β → α → m β) (b l) : by simp [← h (pure b)] induction l with | nil => intro; simp - | cons _ _ l_ih => intro; simp only [List.foldlM, foldl, ←l_ih, functor_norm] + | cons _ _ l_ih => intro; simp only [List.foldlM, foldl, ← l_ih, functor_norm] #align list.mfoldl_eq_foldl List.foldlM_eq_foldl -- Porting note: now in std @@ -2811,8 +2811,8 @@ where | nil => contradiction | cons hd tl => rw [length, succ_eq_add_one] at h - rw [splitAt.go, take, drop, append_cons, Array.toList_eq, ←Array.push_data, - ←Array.toList_eq] + rw [splitAt.go, take, drop, append_cons, Array.toList_eq, ← Array.push_data, + ← Array.toList_eq] exact ih _ _ <| lt_of_add_lt_add_right h · induction n generalizing xs acc with | zero => @@ -2994,7 +2994,7 @@ theorem modifyLast_append (f : α → α) (l₁ l₂ : List α) (_ : l₂ ≠ [] cases tl with | nil => exact modifyLast_append_one _ hd _ | cons hd' tl' => - rw [append_cons, ←nil_append (hd :: hd' :: tl'), append_cons [], nil_append, + rw [append_cons, ← nil_append (hd :: hd' :: tl'), append_cons [], nil_append, modifyLast_append _ (l₁ ++ [hd]) (hd' :: tl') _, modifyLast_append _ [hd] (hd' :: tl') _, append_assoc] all_goals { exact cons_ne_nil _ _ } @@ -3519,8 +3519,8 @@ lemma filter_attach (l : List α) (p : α → Bool) : (l.attach.filter fun x => p x : List {x // x ∈ l}) = (l.filter p).attach.map (Subtype.map id fun x => mem_of_mem_filter) := map_injective_iff.2 Subtype.coe_injective <| by - simp_rw [map_map, (· ∘ ·), Subtype.map, id.def, ←Function.comp_apply (g := Subtype.val), - ←map_filter, attach_map_val] + simp_rw [map_map, (· ∘ ·), Subtype.map, id.def, ← Function.comp_apply (g := Subtype.val), + ← map_filter, attach_map_val] #align list.filter_attach List.filter_attach #align list.filter_filter List.filter_filter @@ -4227,7 +4227,7 @@ theorem forall_iff_forall_mem : ∀ {l : List α}, Forall p l ↔ ∀ x ∈ l, p theorem Forall.imp (h : ∀ x, p x → q x) : ∀ {l : List α}, Forall p l → Forall q l | [] => id - | x :: l => by simp; rw [←and_imp]; exact And.imp (h x) (Forall.imp h) + | x :: l => by simp; rw [← and_imp]; exact And.imp (h x) (Forall.imp h) #align list.all₂.imp List.Forall.imp @[simp] @@ -4310,7 +4310,7 @@ theorem sizeOf_dropSlice_lt [SizeOf α] (i j : ℕ) (hj : 0 < j) (xs : List α) cases n · simp · simp [drop] - rw [←Nat.zero_add (sizeOf (drop _ xs_tl))] + rw [← Nat.zero_add (sizeOf (drop _ xs_tl))] exact Nat.add_le_add (Nat.zero_le _) (drop_sizeOf_le xs_tl _) · simp · simp diff --git a/Mathlib/Data/List/Chain.lean b/Mathlib/Data/List/Chain.lean index 368b6cdf504f9..d70d94cd275ed 100644 --- a/Mathlib/Data/List/Chain.lean +++ b/Mathlib/Data/List/Chain.lean @@ -457,7 +457,7 @@ theorem Chain'.cons_of_le [LinearOrder α] {a : α} {as m : List α} refine gt_of_gt_of_ge ha.1 ?_ rw [le_iff_lt_or_eq] at hmas cases' hmas with hmas hmas - · by_contra' hh + · by_contra! hh rw [← not_le] at hmas apply hmas apply le_of_lt diff --git a/Mathlib/Data/List/Count.lean b/Mathlib/Data/List/Count.lean index 76536cfe8b31d..fb8abd347de34 100644 --- a/Mathlib/Data/List/Count.lean +++ b/Mathlib/Data/List/Count.lean @@ -73,7 +73,7 @@ theorem length_filter_lt_length_iff_exists (l) : -- porting note: `Lean.Internal.coeM` forces us to type-ascript `{x // x ∈ l}` lemma countP_attach (l : List α) : l.attach.countP (fun a : {x // x ∈ l} ↦ p a) = l.countP p := by - simp_rw [←Function.comp_apply (g := Subtype.val), ←countP_map, attach_map_val] + simp_rw [← Function.comp_apply (g := Subtype.val), ← countP_map, attach_map_val] #align list.countp_attach List.countP_attach #align list.countp_mono_left List.countP_mono_left diff --git a/Mathlib/Data/List/Defs.lean b/Mathlib/Data/List/Defs.lean index 77b3d4e67a050..4acee6ce4f8da 100644 --- a/Mathlib/Data/List/Defs.lean +++ b/Mathlib/Data/List/Defs.lean @@ -410,10 +410,11 @@ def chooseX : ∀ l : List α, ∀ _ : ∃ a, a ∈ l ∧ p a, { a // a ∈ l | l :: ls, hp => if pl : p l then ⟨l, ⟨mem_cons.mpr <| Or.inl rfl, pl⟩⟩ else - let ⟨a, ⟨a_mem_ls, pa⟩⟩ := + -- pattern matching on `hx` too makes this not reducible! + let ⟨a, ha⟩ := chooseX ls (hp.imp fun _ ⟨o, h₂⟩ => ⟨(mem_cons.mp o).resolve_left fun e => pl <| e ▸ h₂, h₂⟩) - ⟨a, ⟨mem_cons.mpr <| Or.inr a_mem_ls, pa⟩⟩ + ⟨a, mem_cons.mpr <| Or.inr ha.1, ha.2⟩ #align list.choose_x List.chooseX /-- Given a decidable predicate `p` and a proof of existence of `a ∈ l` such that `p a`, diff --git a/Mathlib/Data/List/EditDistance/Estimator.lean b/Mathlib/Data/List/EditDistance/Estimator.lean index 2da5c124da9d4..6335890f01e14 100644 --- a/Mathlib/Data/List/EditDistance/Estimator.lean +++ b/Mathlib/Data/List/EditDistance/Estimator.lean @@ -88,7 +88,7 @@ instance estimator' : simp only at eq dsimp [EstimatorData.bound] rw [eq] - simp only [←split] + simp only [← split] constructor · simp only [List.minimum_of_length_pos_le_iff] exact suffixLevenshtein_minimum_le_levenshtein_append _ _ _ @@ -112,7 +112,7 @@ instance estimator' : · refine (?_ : _ ≤ _).trans (List.minimum_of_length_pos_le_getElem _) simp only [List.minimum_of_length_pos_le_iff, List.coe_minimum_of_length_pos, d_eq] apply le_suffixLevenshtein_cons_minimum - · simp [←split] + · simp [← split] | y₁ :: y₂ :: t, split, b_eq, d_eq => simp only [EstimatorData.bound, Prod.lt_iff] right diff --git a/Mathlib/Data/List/NodupEquivFin.lean b/Mathlib/Data/List/NodupEquivFin.lean index 9acc1abde58de..34d464b62aa54 100644 --- a/Mathlib/Data/List/NodupEquivFin.lean +++ b/Mathlib/Data/List/NodupEquivFin.lean @@ -129,7 +129,7 @@ theorem sublist_of_orderEmbedding_get?_eq {l l' : List α} (f : ℕ ↪o ℕ) exact b.succ_pos have : ∀ ix, tl.get? ix = (l'.drop (f 0 + 1)).get? (f' ix) := by intro ix - rw [List.get?_drop, OrderEmbedding.coe_ofMapLEIff, add_tsub_cancel_of_le, ←hf, List.get?] + rw [List.get?_drop, OrderEmbedding.coe_ofMapLEIff, add_tsub_cancel_of_le, ← hf, List.get?] rw [Nat.succ_le_iff, OrderEmbedding.lt_iff_lt] exact ix.succ_pos rw [← List.take_append_drop (f 0 + 1) l', ← List.singleton_append] diff --git a/Mathlib/Data/Matrix/Basic.lean b/Mathlib/Data/Matrix/Basic.lean index 4fab3f1f21f1a..3f9c3dcc921b8 100644 --- a/Mathlib/Data/Matrix/Basic.lean +++ b/Mathlib/Data/Matrix/Basic.lean @@ -588,7 +588,7 @@ instance instAddMonoidWithOne [AddMonoidWithOne α] : AddMonoidWithOne (Matrix n natCast_zero := show diagonal _ = _ by rw [Nat.cast_zero, diagonal_zero] natCast_succ n := show diagonal _ = diagonal _ + _ by - rw [Nat.cast_succ, ←diagonal_add, diagonal_one] + rw [Nat.cast_succ, ← diagonal_add, diagonal_one] instance instAddGroupWithOne [AddGroupWithOne α] : AddGroupWithOne (Matrix n n α) where intCast_ofNat n := show diagonal _ = diagonal _ by @@ -1104,12 +1104,12 @@ def addMonoidHomMulRight [Fintype m] (M : Matrix m n α) : Matrix l m α →+ Ma protected theorem sum_mul [Fintype m] (s : Finset β) (f : β → Matrix l m α) (M : Matrix m n α) : (∑ a in s, f a) * M = ∑ a in s, f a * M := - (addMonoidHomMulRight M : Matrix l m α →+ _).map_sum f s + map_sum (addMonoidHomMulRight M) f s #align matrix.sum_mul Matrix.sum_mul protected theorem mul_sum [Fintype m] (s : Finset β) (f : β → Matrix m n α) (M : Matrix l m α) : (M * ∑ a in s, f a) = ∑ a in s, M * f a := - (addMonoidHomMulLeft M : Matrix m n α →+ _).map_sum f s + map_sum (addMonoidHomMulLeft M) f s #align matrix.mul_sum Matrix.mul_sum /-- This instance enables use with `smul_mul_assoc`. -/ @@ -1287,7 +1287,7 @@ theorem scalar_inj [Nonempty n] {r s : α} : scalar n r = scalar n s ↔ r = s : theorem scalar_commute [Fintype n] [DecidableEq n] (r : α) (hr : ∀ r', Commute r r') (M : Matrix n n α) : Commute (scalar n r) M := by - simp_rw [Commute, SemiconjBy, scalar_apply, ←smul_eq_diagonal_mul, ←op_smul_eq_mul_diagonal] + simp_rw [Commute, SemiconjBy, scalar_apply, ← smul_eq_diagonal_mul, ← op_smul_eq_mul_diagonal] ext i j' exact hr _ #align matrix.scalar.commute Matrix.scalar_commuteₓ @@ -2068,7 +2068,7 @@ theorem transpose_multiset_sum [AddCommMonoid α] (s : Multiset (Matrix m n α)) theorem transpose_sum [AddCommMonoid α] {ι : Type*} (s : Finset ι) (M : ι → Matrix m n α) : (∑ i in s, M i)ᵀ = ∑ i in s, (M i)ᵀ := - (transposeAddEquiv m n α).toAddMonoidHom.map_sum _ s + map_sum (transposeAddEquiv m n α) _ s #align matrix.transpose_sum Matrix.transpose_sum variable (m n R α) @@ -2167,7 +2167,7 @@ theorem conjTranspose_zero [AddMonoid α] [StarAddMonoid α] : (0 : Matrix m n @[simp] theorem conjTranspose_eq_zero [AddMonoid α] [StarAddMonoid α] {M : Matrix m n α} : Mᴴ = 0 ↔ M = 0 := - by rw [←conjTranspose_inj (A := M), conjTranspose_zero] + by rw [← conjTranspose_inj (A := M), conjTranspose_zero] @[simp] theorem conjTranspose_one [DecidableEq n] [Semiring α] [StarRing α] : (1 : Matrix n n α)ᴴ = 1 := by @@ -2325,7 +2325,7 @@ theorem conjTranspose_multiset_sum [AddCommMonoid α] [StarAddMonoid α] theorem conjTranspose_sum [AddCommMonoid α] [StarAddMonoid α] {ι : Type*} (s : Finset ι) (M : ι → Matrix m n α) : (∑ i in s, M i)ᴴ = ∑ i in s, (M i)ᴴ := - (conjTransposeAddEquiv m n α).toAddMonoidHom.map_sum _ s + map_sum (conjTransposeAddEquiv m n α) _ s #align matrix.conj_transpose_sum Matrix.conjTranspose_sum variable (m n R α) diff --git a/Mathlib/Data/Matrix/Basis.lean b/Mathlib/Data/Matrix/Basis.lean index 10cdda78715e1..40e9e7a798016 100644 --- a/Mathlib/Data/Matrix/Basis.lean +++ b/Mathlib/Data/Matrix/Basis.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2020 Jalex Stark. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Jalex Stark, Scott Morrison, Eric Wieser, Oliver Nash +Authors: Jalex Stark, Scott Morrison, Eric Wieser, Oliver Nash, Wen Yang -/ import Mathlib.Data.Matrix.Basic import Mathlib.LinearAlgebra.Matrix.Trace @@ -216,6 +216,47 @@ theorem mul_of_ne {k l : n} (h : j ≠ k) (d : α) : end +section Commute + +variable [Fintype n] + +theorem row_eq_zero_of_commute_stdBasisMatrix {i j k : n} {M : Matrix n n α} + (hM : Commute (stdBasisMatrix i j 1) M) (hkj : k ≠ j) : M j k = 0 := by + have := ext_iff.mpr hM i k + aesop + +theorem col_eq_zero_of_commute_stdBasisMatrix {i j k : n} {M : Matrix n n α} + (hM : Commute (stdBasisMatrix i j 1) M) (hki : k ≠ i) : M k i = 0 := by + have := ext_iff.mpr hM k j + aesop + +theorem diag_eq_of_commute_stdBasisMatrix {i j : n} {M : Matrix n n α} + (hM : Commute (stdBasisMatrix i j 1) M) : M i i = M j j := by + have := ext_iff.mpr hM i j + aesop + +/-- `M` is a scalar matrix if it commutes with every non-diagonal `stdBasisMatrix`. ​-/ +theorem mem_range_scalar_of_commute_stdBasisMatrix {M : Matrix n n α} + (hM : ∀ (i j : n), i ≠ j → Commute (stdBasisMatrix i j 1) M) : + M ∈ Set.range (Matrix.scalar n) := by + cases isEmpty_or_nonempty n + · exact ⟨0, Subsingleton.elim _ _⟩ + obtain ⟨i⟩ := ‹Nonempty n› + refine ⟨M i i, Matrix.ext fun j k => ?_⟩ + simp only [scalar_apply] + obtain rfl | hkl := Decidable.eq_or_ne j k + · rw [diagonal_apply_eq] + obtain rfl | hij := Decidable.eq_or_ne i j + · rfl + · exact diag_eq_of_commute_stdBasisMatrix (hM _ _ hij) + · push_neg at hkl + rw [diagonal_apply_ne _ hkl] + obtain rfl | hij := Decidable.eq_or_ne i j + · rw [col_eq_zero_of_commute_stdBasisMatrix (hM k i hkl.symm) hkl] + · rw [row_eq_zero_of_commute_stdBasisMatrix (hM i j hij) hkl.symm] + +end Commute + end StdBasisMatrix end Matrix diff --git a/Mathlib/Data/Matrix/Invertible.lean b/Mathlib/Data/Matrix/Invertible.lean index f5e235f847242..c97640d13814b 100644 --- a/Mathlib/Data/Matrix/Invertible.lean +++ b/Mathlib/Data/Matrix/Invertible.lean @@ -89,8 +89,8 @@ variable [CommSemiring α] (A : Matrix n n α) /-- The transpose of an invertible matrix is invertible. -/ instance invertibleTranspose [Invertible A] : Invertible Aᵀ where invOf := (⅟A)ᵀ - invOf_mul_self := by rw [←transpose_mul, mul_invOf_self, transpose_one] - mul_invOf_self := by rw [←transpose_mul, invOf_mul_self, transpose_one] + invOf_mul_self := by rw [← transpose_mul, mul_invOf_self, transpose_one] + mul_invOf_self := by rw [← transpose_mul, invOf_mul_self, transpose_one] #align matrix.invertible_transpose Matrix.invertibleTranspose lemma transpose_invOf [Invertible A] [Invertible Aᵀ] : (⅟A)ᵀ = ⅟(Aᵀ) := by @@ -100,8 +100,8 @@ lemma transpose_invOf [Invertible A] [Invertible Aᵀ] : (⅟A)ᵀ = ⅟(Aᵀ) : /-- `Aᵀ` is invertible when `A` is. -/ def invertibleOfInvertibleTranspose [Invertible Aᵀ] : Invertible A where invOf := (⅟(Aᵀ))ᵀ - invOf_mul_self := by rw [←transpose_one, ← mul_invOf_self Aᵀ, transpose_mul, transpose_transpose] - mul_invOf_self := by rw [←transpose_one, ← invOf_mul_self Aᵀ, transpose_mul, transpose_transpose] + invOf_mul_self := by rw [← transpose_one, ← mul_invOf_self Aᵀ, transpose_mul, transpose_transpose] + mul_invOf_self := by rw [← transpose_one, ← invOf_mul_self Aᵀ, transpose_mul, transpose_transpose] #align matrix.invertible__of_invertible_transpose Matrix.invertibleOfInvertibleTranspose /-- Together `Matrix.invertibleTranspose` and `Matrix.invertibleOfInvertibleTranspose` form an diff --git a/Mathlib/Data/Matroid/Basic.lean b/Mathlib/Data/Matroid/Basic.lean index 881f588d325d3..3a9d959ef955a 100644 --- a/Mathlib/Data/Matroid/Basic.lean +++ b/Mathlib/Data/Matroid/Basic.lean @@ -253,10 +253,10 @@ theorem encard_diff_le_aux (exch : ExchangeProperty Base) (hB₁ : Base B₁) (h rw [insert_diff_of_mem _ hf.1, diff_diff_comm]; exact hcard have hencard := encard_diff_le_aux exch hB₁ hB' - rw [insert_diff_of_mem _ hf.1, diff_diff_comm, ←union_singleton, ←diff_diff, diff_diff_right, + rw [insert_diff_of_mem _ hf.1, diff_diff_comm, ← union_singleton, ← diff_diff, diff_diff_right, inter_singleton_eq_empty.mpr he.2, union_empty] at hencard - rw [←encard_diff_singleton_add_one he, ←encard_diff_singleton_add_one hf] + rw [← encard_diff_singleton_add_one he, ← encard_diff_singleton_add_one hf] exact add_le_add_right hencard 1 termination_by _ => (B₂ \ B₁).encard @@ -268,7 +268,7 @@ theorem encard_diff_eq (hB₁ : Base B₁) (hB₂ : Base B₂) : (B₁ \ B₂).e /-- Any two sets `B₁`, `B₂` in a family with the exchange property have the same `ℕ∞`-cardinality. -/ theorem encard_base_eq (hB₁ : Base B₁) (hB₂ : Base B₂) : B₁.encard = B₂.encard := by -rw [←encard_diff_add_encard_inter B₁ B₂, exch.encard_diff_eq hB₁ hB₂, inter_comm, +rw [← encard_diff_add_encard_inter B₁ B₂, exch.encard_diff_eq hB₁ hB₂, inter_comm, encard_diff_add_encard_inter] end ExchangeProperty @@ -357,14 +357,14 @@ theorem Base.encard_diff_comm (hB₁ : M.Base B₁) (hB₂ : M.Base B₂) : theorem Base.ncard_diff_comm (hB₁ : M.Base B₁) (hB₂ : M.Base B₂) : (B₁ \ B₂).ncard = (B₂ \ B₁).ncard := by - rw [ncard_def, hB₁.encard_diff_comm hB₂, ←ncard_def] + rw [ncard_def, hB₁.encard_diff_comm hB₂, ← ncard_def] theorem Base.card_eq_card_of_base (hB₁ : M.Base B₁) (hB₂ : M.Base B₂) : B₁.encard = B₂.encard := by rw [M.base_exchange.encard_base_eq hB₁ hB₂] theorem Base.ncard_eq_ncard_of_base (hB₁ : M.Base B₁) (hB₂ : M.Base B₂) : B₁.ncard = B₂.ncard := by - rw [ncard_def B₁, hB₁.card_eq_card_of_base hB₂, ←ncard_def] + rw [ncard_def B₁, hB₁.card_eq_card_of_base hB₂, ← ncard_def] theorem Base.finite_of_finite (hB : M.Base B) (h : B.Finite) (hB' : M.Base B') : B'.Finite := (finite_iff_finite_of_encard_eq_encard (hB.card_eq_card_of_base hB')).mp h @@ -430,7 +430,7 @@ theorem base_compl_iff_mem_maximals_disjoint_base (hB : B ⊆ M.E := by aesop_ma simp_rw [mem_maximals_setOf_iff, and_iff_right hB, and_imp, forall_exists_index] refine' ⟨fun h ↦ ⟨⟨_, h, disjoint_sdiff_right⟩, fun I hI B' ⟨hB', hIB'⟩ hBI ↦ hBI.antisymm _⟩, fun ⟨⟨ B', hB', hBB'⟩,h⟩ ↦ _⟩ - · rw [hB'.eq_of_subset_base h, ←subset_compl_iff_disjoint_right, diff_eq, compl_inter, + · rw [hB'.eq_of_subset_base h, ← subset_compl_iff_disjoint_right, diff_eq, compl_inter, compl_compl] at hIB' · exact fun e he ↦ (hIB' he).elim (fun h' ↦ (h' (hI he)).elim) id rw [subset_diff, and_iff_right hB'.subset_ground, disjoint_comm] @@ -558,14 +558,14 @@ theorem Base.eq_exchange_of_diff_eq_singleton (hB : M.Base B) (hB' : M.Base B') have hne : f ≠ e := by rintro rfl; exact hf.2 (h.symm.subset (mem_singleton f)).1 rw [insert_diff_singleton_comm hne] at hb refine ⟨f, hf, (hb.eq_of_subset_base hB' ?_).symm⟩ - rw [diff_subset_iff, insert_subset_iff, union_comm, ←diff_subset_iff, h, and_iff_left rfl.subset] + rw [diff_subset_iff, insert_subset_iff, union_comm, ← diff_subset_iff, h, and_iff_left rfl.subset] exact Or.inl hf.1 theorem Base.exchange_base_of_indep (hB : M.Base B) (hf : f ∉ B) (hI : M.Indep (insert f (B \ {e}))) : M.Base (insert f (B \ {e})) := by obtain ⟨B', hB', hIB'⟩ := hI.exists_base_superset have hcard := hB'.encard_diff_comm hB - rw [insert_subset_iff, ←diff_eq_empty, diff_diff_comm, diff_eq_empty, subset_singleton_iff_eq] + rw [insert_subset_iff, ← diff_eq_empty, diff_diff_comm, diff_eq_empty, subset_singleton_iff_eq] at hIB' obtain ⟨hfB, (h | h)⟩ := hIB' · rw [h, encard_empty, encard_eq_zero, eq_empty_iff_forall_not_mem] at hcard @@ -573,18 +573,18 @@ theorem Base.exchange_base_of_indep (hB : M.Base B) (hf : f ∉ B) rw [h, encard_singleton, encard_eq_one] at hcard obtain ⟨x, hx⟩ := hcard obtain (rfl : f = x) := hx.subset ⟨hfB, hf⟩ - simp_rw [←h, ←singleton_union, ←hx, sdiff_sdiff_right_self, inf_eq_inter, inter_comm B, + simp_rw [← h, ← singleton_union, ← hx, sdiff_sdiff_right_self, inf_eq_inter, inter_comm B, diff_union_inter] exact hB' theorem Base.exchange_base_of_indep' (hB : M.Base B) (he : e ∈ B) (hf : f ∉ B) (hI : M.Indep (insert f B \ {e})) : M.Base (insert f B \ {e}) := by have hfe : f ≠ e := by rintro rfl; exact hf he - rw [←insert_diff_singleton_comm hfe] at * + rw [← insert_diff_singleton_comm hfe] at * exact hB.exchange_base_of_indep hf hI theorem Base.insert_dep (hB : M.Base B) (h : e ∈ M.E \ B) : M.Dep (insert e B) := by - rw [←not_indep_iff (insert_subset h.1 hB.subset_ground)] + rw [← not_indep_iff (insert_subset h.1 hB.subset_ground)] exact h.2 ∘ (fun hi ↦ insert_eq_self.mp (hB.eq_of_subset_indep hi (subset_insert e B)).symm) theorem Indep.exists_insert_of_not_base (hI : M.Indep I) (hI' : ¬M.Base I) (hB : M.Base B) : @@ -774,7 +774,7 @@ theorem Indep.basis_self (h : M.Indep I) : M.Basis I I := theorem Basis.dep_of_ssubset (hI : M.Basis I X) (hIY : I ⊂ Y) (hYX : Y ⊆ X) : M.Dep Y := by have : X ⊆ M.E := hI.subset_ground - rw [←not_indep_iff] + rw [← not_indep_iff] exact fun hY ↦ hIY.ne (hI.eq_of_subset_indep hY hIY.subset hYX) theorem Basis.insert_dep (hI : M.Basis I X) (he : e ∈ X \ I) : M.Dep (insert e I) := @@ -851,7 +851,7 @@ theorem Indep.basis_iff_forall_insert_dep (hI : M.Indep I) (hIX : I ⊆ X) : fun h ↦ ⟨fun J hJ hIJ hJX ↦ hIJ.antisymm (fun e heJ ↦ by_contra (fun heI ↦ _)),_⟩⟩ · exact (h.1 _ hi (subset_insert _ _) (insert_subset he.1 hIX)).symm.subset (mem_insert e I) · exact (h e ⟨hJX heJ, heI⟩).not_indep (hJ.subset (insert_subset heJ hIJ)) - rw [←diff_union_of_subset hIX, union_subset_iff, and_iff_left hI.subset_ground] + rw [← diff_union_of_subset hIX, union_subset_iff, and_iff_left hI.subset_ground] exact fun e he ↦ (h e he).subset_ground (mem_insert _ _) theorem Indep.basis_of_forall_insert (hI : M.Indep I) (hIX : I ⊆ X) @@ -898,7 +898,7 @@ theorem Basis.union_basis_union (hIX : M.Basis I X) (hJY : M.Basis J Y) (h : M.I rw [union_eq_iUnion, union_eq_iUnion] refine' Basis.iUnion_basis_iUnion _ _ _ _ · simp only [Bool.forall_bool, cond_false, cond_true]; exact ⟨hJY, hIX⟩ - rwa [←union_eq_iUnion] + rwa [← union_eq_iUnion] theorem Basis.basis_union (hIX : M.Basis I X) (hIY : M.Basis I Y) : M.Basis I (X ∪ Y) := by convert hIX.union_basis_union hIY _ <;> rw [union_self]; exact hIX.indep @@ -911,7 +911,7 @@ theorem Basis.basis_union_of_subset (hI : M.Basis I X) (hJ : M.Indep J) (hIJ : I theorem Basis.insert_basis_insert (hI : M.Basis I X) (h : M.Indep (insert e I)) : M.Basis (insert e I) (insert e X) := by - simp_rw [←union_singleton] at * + simp_rw [← union_singleton] at * exact hI.union_basis_union (h.subset (subset_union_right _ _)).basis_self h theorem Base.base_of_basis_superset (hB : M.Base B) (hBX : B ⊆ X) (hIX : M.Basis I X) : diff --git a/Mathlib/Data/Matroid/IndepAxioms.lean b/Mathlib/Data/Matroid/IndepAxioms.lean index 7e0b1a47b7dcf..cfb84204a442d 100644 --- a/Mathlib/Data/Matroid/IndepAxioms.lean +++ b/Mathlib/Data/Matroid/IndepAxioms.lean @@ -236,7 +236,7 @@ attribute [pp_dot] Indep E ⟨hB₀J.trans <| subset_insert _ _,hfi,insert_subset hfE₀ hJss⟩ (subset_insert _ _))) -- But this means `|I₀| < |J|`, and extending `I₀` into `J` gives a contradiction - rw [ncard_insert_of_not_mem heI₀ hI₀fin, ←Nat.lt_iff_add_one_le] at hcard + rw [ncard_insert_of_not_mem heI₀ hI₀fin, ← Nat.lt_iff_add_one_le] at hcard obtain ⟨f, hfJ, hfI₀, hfi⟩ := indep_aug (indep_subset hI hI₀I) hI₀fin hJ hJfin hcard exact hI₀ f ⟨Or.elim (hJss hfJ) (fun hfe ↦ (heJ <| hfe ▸ hfJ).elim) (by aesop), hfI₀⟩ hfi ) @@ -378,7 +378,7 @@ protected def ofFinite {E : Set α} (hE : E.Finite) (Indep : Set α → Prop) (indep_subset := indep_subset) (indep_aug := by refine fun {I J} hI hJ hIJ ↦ indep_aug hI hJ ?_ - rwa [←Nat.cast_lt (α := ℕ∞), (hE.subset (subset_ground hJ)).cast_ncard_eq, + rwa [← Nat.cast_lt (α := ℕ∞), (hE.subset (subset_ground hJ)).cast_ncard_eq, (hE.subset (subset_ground hI)).cast_ncard_eq] ) (indep_bdd := ⟨E.ncard, fun I hI ↦ by rw [hE.cast_ncard_eq] diff --git a/Mathlib/Data/Multiset/Basic.lean b/Mathlib/Data/Multiset/Basic.lean index a6960ea6cdef6..4a7e31047c627 100644 --- a/Mathlib/Data/Multiset/Basic.lean +++ b/Mathlib/Data/Multiset/Basic.lean @@ -2487,25 +2487,25 @@ theorem count_eq_card {a : α} {s} : count a s = card s ↔ ∀ x ∈ s, a = x : @[simp] theorem count_replicate_self (a : α) (n : ℕ) : count a (replicate n a) = n := by convert List.count_replicate_self a n - rw [←coe_count, coe_replicate] + rw [← coe_count, coe_replicate] #align multiset.count_replicate_self Multiset.count_replicate_self theorem count_replicate (a b : α) (n : ℕ) : count a (replicate n b) = if a = b then n else 0 := by convert List.count_replicate a b n - rw [←coe_count, coe_replicate] + rw [← coe_count, coe_replicate] #align multiset.count_replicate Multiset.count_replicate @[simp] theorem count_erase_self (a : α) (s : Multiset α) : count a (erase s a) = count a s - 1 := Quotient.inductionOn s <| fun l => by - convert List.count_erase_self a l <;> rw [←coe_count] <;> simp + convert List.count_erase_self a l <;> rw [← coe_count] <;> simp #align multiset.count_erase_self Multiset.count_erase_self @[simp] theorem count_erase_of_ne {a b : α} (ab : a ≠ b) (s : Multiset α) : count a (erase s b) = count a s := Quotient.inductionOn s <| fun l => by - convert List.count_erase_of_ne ab l <;> rw [←coe_count] <;> simp + convert List.count_erase_of_ne ab l <;> rw [← coe_count] <;> simp #align multiset.count_erase_of_ne Multiset.count_erase_of_ne @[simp] diff --git a/Mathlib/Data/Multiset/Bind.lean b/Mathlib/Data/Multiset/Bind.lean index 3aa8a6c1d6891..c4fc8b85216ba 100644 --- a/Mathlib/Data/Multiset/Bind.lean +++ b/Mathlib/Data/Multiset/Bind.lean @@ -41,7 +41,7 @@ theorem coe_join : | [] => rfl | l :: L => by -- Porting note: was `congr_arg (fun s : Multiset α => ↑l + s) (coe_join L)` - simp only [join, List.map, coe_sum, List.sum_cons, List.join, ←coe_add, ←coe_join L] + simp only [join, List.map, coe_sum, List.sum_cons, List.join, ← coe_add, ← coe_join L] #align multiset.coe_join Multiset.coe_join @[simp] diff --git a/Mathlib/Data/Multiset/Fintype.lean b/Mathlib/Data/Multiset/Fintype.lean index 9173734a61b59..a6b7a28c423ab 100644 --- a/Mathlib/Data/Multiset/Fintype.lean +++ b/Mathlib/Data/Multiset/Fintype.lean @@ -108,7 +108,7 @@ instance : Fintype { p : α × ℕ | p.2 < m.count p.1 } := rintro ⟨x, i⟩ simp only [Finset.mem_biUnion, Multiset.mem_toFinset, Finset.mem_map, Finset.mem_range, Function.Embedding.coeFn_mk, Prod.mk.inj_iff, Set.mem_setOf_eq] - simp only [←and_assoc, exists_eq_right, and_iff_right_iff_imp] + simp only [← and_assoc, exists_eq_right, and_iff_right_iff_imp] exact fun h ↦ Multiset.count_pos.mp (pos_of_gt h)) /-- Construct a finset whose elements enumerate the elements of the multiset `m`. @@ -245,7 +245,7 @@ theorem Multiset.map_univ {β : Type*} (m : Multiset α) (f : α → β) : @[simp] theorem Multiset.card_toEnumFinset (m : Multiset α) : m.toEnumFinset.card = Multiset.card m := by - rw [Finset.card, ←Multiset.card_map Prod.fst m.toEnumFinset.val] + rw [Finset.card, ← Multiset.card_map Prod.fst m.toEnumFinset.val] congr exact m.map_toEnumFinset_fst #align multiset.card_to_enum_finset Multiset.card_toEnumFinset diff --git a/Mathlib/Data/Multiset/Nodup.lean b/Mathlib/Data/Multiset/Nodup.lean index 1362ec84fe0cc..0e7e6e9063c5a 100644 --- a/Mathlib/Data/Multiset/Nodup.lean +++ b/Mathlib/Data/Multiset/Nodup.lean @@ -203,7 +203,7 @@ protected theorem Nodup.sigma {σ : α → Type*} {t : ∀ a, Multiset (σ a)} : Nodup s → (∀ a, Nodup (t a)) → Nodup (s.sigma t) := Quot.induction_on s fun l₁ => by choose f hf using fun a => Quotient.exists_rep (t a) - simpa [←funext hf] using List.Nodup.sigma + simpa [← funext hf] using List.Nodup.sigma #align multiset.nodup.sigma Multiset.Nodup.sigma protected theorem Nodup.filterMap (f : α → Option β) (H : ∀ a a' b, b ∈ f a → b ∈ f a' → a = a') : diff --git a/Mathlib/Data/MvPolynomial/Basic.lean b/Mathlib/Data/MvPolynomial/Basic.lean index 784a7c62f4d87..6120acc00166c 100644 --- a/Mathlib/Data/MvPolynomial/Basic.lean +++ b/Mathlib/Data/MvPolynomial/Basic.lean @@ -247,7 +247,7 @@ theorem C_injective (σ : Type*) (R : Type*) [CommSemiring R] : theorem C_surjective {R : Type*} [CommSemiring R] (σ : Type*) [IsEmpty σ] : Function.Surjective (C : R → MvPolynomial σ R) := by refine' fun p => ⟨p.toFun 0, Finsupp.ext fun a => _⟩ - simp only [C_apply, ←single_eq_monomial, (Finsupp.ext isEmptyElim (α := σ) : a = 0), + simp only [C_apply, ← single_eq_monomial, (Finsupp.ext isEmptyElim (α := σ) : a = 0), single_eq_same] rfl #align mv_polynomial.C_surjective MvPolynomial.C_surjective @@ -373,7 +373,7 @@ theorem sum_C {A : Type*} [AddCommMonoid A] {b : (σ →₀ ℕ) → R → A} (w theorem monomial_sum_one {α : Type*} (s : Finset α) (f : α → σ →₀ ℕ) : (monomial (∑ i in s, f i) 1 : MvPolynomial σ R) = ∏ i in s, monomial (f i) 1 := - (monomialOneHom R σ).map_prod (fun i => Multiplicative.ofAdd (f i)) s + map_prod (monomialOneHom R σ) (fun i => Multiplicative.ofAdd (f i)) s #align mv_polynomial.monomial_sum_one MvPolynomial.monomial_sum_one theorem monomial_sum_index {α : Type*} (s : Finset α) (f : α → σ →₀ ℕ) (a : R) : @@ -533,7 +533,7 @@ theorem finsupp_support_eq_support (p : MvPolynomial σ R) : Finsupp.support p = theorem support_monomial [h : Decidable (a = 0)] : (monomial s a).support = if a = 0 then ∅ else {s} := by - rw [←Subsingleton.elim (Classical.decEq R a 0) h] + rw [← Subsingleton.elim (Classical.decEq R a 0) h] rfl -- porting note: the proof in Lean 3 wasn't fundamentally better and needed `by convert rfl` -- the issue is the different decidability instances in the `ite` expressions @@ -642,7 +642,7 @@ def coeffAddMonoidHom (m : σ →₀ ℕ) : MvPolynomial σ R →+ R theorem coeff_sum {X : Type*} (s : Finset X) (f : X → MvPolynomial σ R) (m : σ →₀ ℕ) : coeff m (∑ x in s, f x) = ∑ x in s, coeff m (f x) := - (@coeffAddMonoidHom R σ _ _).map_sum _ s + map_sum (@coeffAddMonoidHom R σ _ _) _ s #align mv_polynomial.coeff_sum MvPolynomial.coeff_sum theorem monic_monomial_eq (m) : @@ -670,7 +670,7 @@ theorem coeff_one [DecidableEq σ] (m) : coeff m (1 : MvPolynomial σ R) = if 0 coeff_C m 1 #align mv_polynomial.coeff_one MvPolynomial.coeff_one --- porting note: `simp` can prove this +@[simp] theorem coeff_zero_C (a) : coeff 0 (C a : MvPolynomial σ R) = a := single_eq_same #align mv_polynomial.coeff_zero_C MvPolynomial.coeff_zero_C diff --git a/Mathlib/Data/MvPolynomial/CommRing.lean b/Mathlib/Data/MvPolynomial/CommRing.lean index e59bef54bdc09..fcb0b40bb546b 100644 --- a/Mathlib/Data/MvPolynomial/CommRing.lean +++ b/Mathlib/Data/MvPolynomial/CommRing.lean @@ -189,7 +189,7 @@ theorem degreeOf_sub_lt {x : σ} {f g : MvPolynomial σ R} {k : ℕ} (h : 0 < k) classical rw [degreeOf_lt_iff h] intro m hm - by_contra' hc + by_contra! hc have h := support_sub σ f g hm simp only [mem_support_iff, Ne.def, coeff_sub, sub_eq_zero] at hm cases' Finset.mem_union.1 h with cf cg diff --git a/Mathlib/Data/MvPolynomial/Equiv.lean b/Mathlib/Data/MvPolynomial/Equiv.lean index 65975d9d0ee8f..7349c6deddb07 100644 --- a/Mathlib/Data/MvPolynomial/Equiv.lean +++ b/Mathlib/Data/MvPolynomial/Equiv.lean @@ -439,7 +439,7 @@ lemma totalDegree_coeff_finSuccEquiv_add_le (f : MvPolynomial (Fin (n + 1)) R) ( let σ' : Fin (n+1) →₀ ℕ := cons i σ convert le_totalDegree (s := σ') _ · rw [totalDegree, hσ2, sum_cons, add_comm] - · rw [←support_coeff_finSuccEquiv] + · rw [← support_coeff_finSuccEquiv] exact hσ1 theorem finSuccEquiv_support (f : MvPolynomial (Fin (n + 1)) R) : diff --git a/Mathlib/Data/Nat/Basic.lean b/Mathlib/Data/Nat/Basic.lean index 5f999e12a50f9..26d9f8dd4fb27 100644 --- a/Mathlib/Data/Nat/Basic.lean +++ b/Mathlib/Data/Nat/Basic.lean @@ -716,7 +716,7 @@ theorem div_mul_div_comm {l : ℕ} (hmn : n ∣ m) (hkl : l ∣ k) : rcases l.eq_zero_or_pos with rfl | hl · simp rw [Nat.mul_div_cancel_left _ hn, Nat.mul_div_cancel_left _ hl, mul_assoc n, Nat.mul_left_comm x, - ←mul_assoc n, Nat.mul_div_cancel_left _ (Nat.mul_pos hn hl)] + ← mul_assoc n, Nat.mul_div_cancel_left _ (Nat.mul_pos hn hl)] #align nat.div_mul_div_comm Nat.div_mul_div_comm protected theorem div_pow {a b c : ℕ} (h : a ∣ b) : (b / a) ^ c = b ^ c / a ^ c := by @@ -725,7 +725,7 @@ protected theorem div_pow {a b c : ℕ} (h : a ∣ b) : (b / a) ^ c = b ^ c / a rcases a.eq_zero_or_pos with rfl | ha · simp [Nat.zero_pow hc] refine (Nat.div_eq_of_eq_mul_right (pos_pow_of_pos c ha) ?_).symm - rw [←Nat.mul_pow, Nat.mul_div_cancel_left' h] + rw [← Nat.mul_pow, Nat.mul_div_cancel_left' h] /-! ### `mod`, `dvd` -/ @@ -811,7 +811,7 @@ theorem mul_dvd_of_dvd_div {a b c : ℕ} (hab : c ∣ b) (h : a ∣ b / c) : c * let ⟨d, hd⟩ := h1 have h3 : b = a * d * c := Nat.eq_mul_of_div_eq_left hab hd -- Porting note: was `cc` - show ∃ d, b = c * a * d from ⟨d, by rwa [mul_comm, ←mul_assoc] at h3⟩ + show ∃ d, b = c * a * d from ⟨d, by rwa [mul_comm, ← mul_assoc] at h3⟩ #align nat.mul_dvd_of_dvd_div Nat.mul_dvd_of_dvd_div theorem eq_of_dvd_of_div_eq_one {a b : ℕ} (w : a ∣ b) (h : b / a = 1) : a = b := by @@ -838,11 +838,12 @@ theorem lt_mul_div_succ (m : ℕ) {n : ℕ} (n0 : 0 < n) : m < n * (m / n + 1) : exact lt_succ_self _ #align nat.lt_mul_div_succ Nat.lt_mul_div_succ -theorem mul_add_mod (a b c : ℕ) : (a * b + c) % b = c % b := by simp [Nat.add_mod] -#align nat.mul_add_mod Nat.mul_add_mod +-- TODO: Std4 claimed this name but flipped the order of multiplication +theorem mul_add_mod' (a b c : ℕ) : (a * b + c) % b = c % b := by rw [mul_comm, Nat.mul_add_mod] +#align nat.mul_add_mod Nat.mul_add_mod' theorem mul_add_mod_of_lt {a b c : ℕ} (h : c < b) : (a * b + c) % b = c := by - rw [Nat.mul_add_mod, Nat.mod_eq_of_lt h] + rw [Nat.mul_add_mod', Nat.mod_eq_of_lt h] #align nat.mul_add_mod_of_lt Nat.mul_add_mod_of_lt theorem pred_eq_self_iff {n : ℕ} : n.pred = n ↔ n = 0 := by diff --git a/Mathlib/Data/Nat/Bitwise.lean b/Mathlib/Data/Nat/Bitwise.lean index 719464d7c6a94..c3b2e42266895 100644 --- a/Mathlib/Data/Nat/Bitwise.lean +++ b/Mathlib/Data/Nat/Bitwise.lean @@ -92,7 +92,7 @@ lemma bitwise_bit {f : Bool → Bool → Bool} (h : f false false = false := by lemma bit_mod_two (a : Bool) (x : ℕ) : bit a x % 2 = if a then 1 else 0 := by - simp (config := { unfoldPartialApp := true }) [bit, bit0, bit1, Bool.cond_eq_ite, ←mul_two] + simp (config := { unfoldPartialApp := true }) [bit, bit0, bit1, Bool.cond_eq_ite, ← mul_two] split_ifs <;> simp [Nat.add_mod] @[simp] @@ -125,15 +125,7 @@ theorem xor_bit : ∀ a m b n, bit a m ^^^ bit b n = bit (bne a b) (m ^^^ n) := bitwise_bit #align nat.lxor_bit Nat.xor_bit -@[simp] -theorem testBit_bitwise {f : Bool → Bool → Bool} (h : f false false = false) (m n k) : - testBit (bitwise f m n) k = f (testBit m k) (testBit n k) := by - induction' k with k ih generalizing m n - <;> cases' m using bitCasesOn with a m - <;> cases' n using bitCasesOn with b n - <;> rw [bitwise_bit h] - · simp [testBit_zero] - · simp [testBit_succ, ih] +attribute [simp] Nat.testBit_bitwise #align nat.test_bit_bitwise Nat.testBit_bitwise @[simp] @@ -151,9 +143,7 @@ theorem testBit_ldiff : ∀ m n k, testBit (ldiff m n) k = (testBit m k && not ( testBit_bitwise rfl #align nat.test_bit_ldiff Nat.testBit_ldiff -@[simp] -theorem testBit_xor : ∀ m n k, testBit (m ^^^ n) k = bne (testBit m k) (testBit n k) := - testBit_bitwise rfl +attribute [simp] testBit_xor #align nat.test_bit_lxor Nat.testBit_xor end @@ -183,7 +173,7 @@ lemma bitwise_bit' {f : Bool → Bool → Bool} (a : Bool) (m : Nat) (b : Bool) (ham : m = 0 → a = true) (hbn : n = 0 → b = true) : bitwise f (bit a m) (bit b n) = bit (f a b) (bitwise f m n) := by conv_lhs => unfold bitwise - rw [←bit_ne_zero_iff] at ham hbn + rw [← bit_ne_zero_iff] at ham hbn simp only [ham, hbn, bit_mod_two_eq_one_iff, Bool.decide_coe, ← div2_val, div2_bit, ne_eq, ite_false] conv_rhs => simp only [bit, bit1, bit0, Bool.cond_eq_ite] @@ -197,13 +187,13 @@ lemma bitwise_eq_binaryRec (f : Bool → Bool → Bool) : induction x using binaryRec' generalizing y with | z => simp only [bitwise_zero_left, binaryRec_zero, Bool.cond_eq_ite] | f xb x hxb ih => - rw [←bit_ne_zero_iff] at hxb + rw [← bit_ne_zero_iff] at hxb simp_rw [binaryRec_of_ne_zero _ _ hxb, bodd_bit, div2_bit, eq_rec_constant] induction y using binaryRec' with | z => simp only [bitwise_zero_right, binaryRec_zero, Bool.cond_eq_ite] | f yb y hyb => - rw [←bit_ne_zero_iff] at hyb - simp_rw [binaryRec_of_ne_zero _ _ hyb, bitwise_of_ne_zero hxb hyb, bodd_bit, ←div2_val, + rw [← bit_ne_zero_iff] at hyb + simp_rw [binaryRec_of_ne_zero _ _ hyb, bitwise_of_ne_zero hxb hyb, bodd_bit, ← div2_val, div2_bit, eq_rec_constant, ih] theorem zero_of_testBit_eq_false {n : ℕ} (h : ∀ i, testBit n i = false) : n = 0 := by @@ -216,31 +206,19 @@ theorem zero_of_testBit_eq_false {n : ℕ} (h : ∀ i, testBit n i = false) : n theorem testBit_eq_false_of_lt {n i} (h : n < 2 ^ i) : n.testBit i = false := by simp [testBit, shiftRight_eq_div_pow, Nat.div_eq_of_lt h] -@[simp] -theorem zero_testBit (i : ℕ) : testBit 0 i = false := by - simp only [testBit, zero_shiftRight, bodd_zero] #align nat.zero_test_bit Nat.zero_testBit /-- The ith bit is the ith element of `n.bits`. -/ theorem testBit_eq_inth (n i : ℕ) : n.testBit i = n.bits.getI i := by induction' i with i ih generalizing n - · simp [testBit, bodd_eq_bits_head, List.getI_zero_eq_headI] + · simp only [testBit, zero_eq, shiftRight_zero, and_one_is_mod, mod_two_of_bodd, + bodd_eq_bits_head, List.getI_zero_eq_headI] + cases List.headI (bits n) <;> rfl conv_lhs => rw [← bit_decomp n] rw [testBit_succ, ih n.div2, div2_bits_eq_tail] cases n.bits <;> simp #align nat.test_bit_eq_inth Nat.testBit_eq_inth -/-- Bitwise extensionality: Two numbers agree if they agree at every bit position. -/ -theorem eq_of_testBit_eq {n m : ℕ} (h : ∀ i, testBit n i = testBit m i) : n = m := by - induction' n using Nat.binaryRec with b n hn generalizing m - · simp only [zero_testBit] at h - exact (zero_of_testBit_eq_false fun i => (h i).symm).symm - induction' m using Nat.binaryRec with b' m - · simp only [zero_testBit] at h - exact zero_of_testBit_eq_false h - suffices h' : n = m by - rw [h', show b = b' by simpa using h 0] - exact hn fun i => by convert h (i + 1) using 1 <;> rw [testBit_succ] #align nat.eq_of_test_bit_eq Nat.eq_of_testBit_eq theorem exists_most_significant_bit {n : ℕ} (h : n ≠ 0) : @@ -294,20 +272,20 @@ theorem lt_of_testBit {n m : ℕ} (i : ℕ) (hn : testBit n i = false) (hm : tes @[simp] theorem testBit_two_pow_self (n : ℕ) : testBit (2 ^ n) n = true := by - rw [testBit, shiftRight_eq_div_pow, Nat.div_self (pow_pos (α := ℕ) zero_lt_two n), bodd_one] + rw [testBit, shiftRight_eq_div_pow, Nat.div_self (pow_pos (α := ℕ) zero_lt_two n)] + simp #align nat.test_bit_two_pow_self Nat.testBit_two_pow_self theorem testBit_two_pow_of_ne {n m : ℕ} (hm : n ≠ m) : testBit (2 ^ n) m = false := by rw [testBit, shiftRight_eq_div_pow] cases' hm.lt_or_lt with hm hm - · rw [Nat.div_eq_of_lt, bodd_zero] - exact Nat.pow_lt_pow_of_lt_right one_lt_two hm + · rw [Nat.div_eq_of_lt] + · simp + · exact Nat.pow_lt_pow_of_lt_right one_lt_two hm · rw [pow_div hm.le zero_lt_two, ← tsub_add_cancel_of_le (succ_le_of_lt <| tsub_pos_of_lt hm)] -- Porting note: XXX why does this make it work? rw [(rfl : succ 0 = 1)] - simp only [ge_iff_le, tsub_le_iff_right, pow_succ, bodd_mul, - Bool.and_eq_false_eq_eq_false_or_eq_false, or_true] - exact Or.inr rfl + simp [pow_succ, and_one_is_mod, mul_mod_left] #align nat.test_bit_two_pow_of_ne Nat.testBit_two_pow_of_ne theorem testBit_two_pow (n m : ℕ) : testBit (2 ^ n) m = (n = m) := by @@ -371,25 +349,10 @@ theorem zero_xor (n : ℕ) : 0 ^^^ n = n := by simp [HXor.hXor, Xor.xor, xor] theorem xor_zero (n : ℕ) : n ^^^ 0 = n := by simp [HXor.hXor, Xor.xor, xor] #align nat.lxor_zero Nat.xor_zero -@[simp] -theorem zero_land (n : ℕ) : 0 &&& n = 0 := by - simp only [HAnd.hAnd, AndOp.and, land, bitwise_zero_left, ite_false, Bool.and_true]; -#align nat.zero_land Nat.zero_land - -@[simp] -theorem land_zero (n : ℕ) : n &&& 0 = 0 := by - simp only [HAnd.hAnd, AndOp.and, land, bitwise_zero_right, ite_false, Bool.and_false] -#align nat.land_zero Nat.land_zero - -@[simp] -theorem zero_lor (n : ℕ) : 0 ||| n = n := by - simp only [HOr.hOr, OrOp.or, lor, bitwise_zero_left, ite_true, Bool.or_true] -#align nat.zero_lor Nat.zero_lor - -@[simp] -theorem lor_zero (n : ℕ) : n ||| 0 = n := by - simp only [HOr.hOr, OrOp.or, lor, bitwise_zero_right, ite_true, Bool.or_false] -#align nat.lor_zero Nat.lor_zero +#align nat.zero_land Nat.zero_and +#align nat.land_zero Nat.and_zero +#align nat.zero_lor Nat.zero_or +#align nat.lor_zero Nat.or_zero /-- Proving associativity of bitwise operations in general essentially boils down to a huge case @@ -487,11 +450,11 @@ theorem xor_trichotomy {a b c : ℕ} (h : a ≠ b ^^^ c) : all_goals exact lt_of_testBit i (by -- Porting note: this was originally `simp [h, hi]` - rw [Nat.testBit_xor, h, Bool.bne_eq_xor, Bool.true_xor,hi] + rw [Nat.testBit_xor, h, Bool.xor, Bool.true_xor, hi] rfl ) h fun j hj => by -- Porting note: this was originally `simp [hi' _ hj]` - rw [Nat.testBit_xor, hi' _ hj, Bool.bne_eq_xor, Bool.xor_false, eq_self_iff_true] + rw [Nat.testBit_xor, hi' _ hj, Bool.xor, Bool.xor_false, eq_self_iff_true] trivial #align nat.lxor_trichotomy Nat.xor_trichotomy @@ -500,7 +463,7 @@ theorem lt_xor_cases {a b c : ℕ} (h : a < b ^^^ c) : a ^^^ c < b ∨ a ^^^ b < #align nat.lt_lxor_cases Nat.lt_xor_cases @[simp] theorem bit_lt_two_pow_succ_iff {b x n} : bit b x < 2 ^ (n + 1) ↔ x < 2 ^ n := by - rw [pow_succ', ←bit0_eq_two_mul] + rw [pow_succ', ← bit0_eq_two_mul] cases b <;> simp /-- If `x` and `y` fit within `n` bits, then the result of any bitwise operation on `x` and `y` also diff --git a/Mathlib/Data/Nat/Cast/Defs.lean b/Mathlib/Data/Nat/Cast/Defs.lean index b408bd87e8e32..0d3c9267ccdb0 100644 --- a/Mathlib/Data/Nat/Cast/Defs.lean +++ b/Mathlib/Data/Nat/Cast/Defs.lean @@ -171,11 +171,11 @@ theorem binCast_eq [AddMonoidWithOne R] (n : ℕ) : (Nat.binCast n : R) = ((n : | succ k => rw [Nat.binCast] by_cases h : (k + 1) % 2 = 0 - · rw [←Nat.mod_add_div (succ k) 2] - rw [if_pos h, hk _ $ Nat.div_lt_self (Nat.succ_pos k) (Nat.le_refl 2), ←Nat.cast_add] + · rw [← Nat.mod_add_div (succ k) 2] + rw [if_pos h, hk _ $ Nat.div_lt_self (Nat.succ_pos k) (Nat.le_refl 2), ← Nat.cast_add] rw [Nat.succ_eq_add_one, h, Nat.zero_add, Nat.succ_mul, Nat.one_mul] - · rw [←Nat.mod_add_div (succ k) 2] - rw [if_neg h, hk _ $ Nat.div_lt_self (Nat.succ_pos k) (Nat.le_refl 2), ←Nat.cast_add] + · rw [← Nat.mod_add_div (succ k) 2] + rw [if_neg h, hk _ $ Nat.div_lt_self (Nat.succ_pos k) (Nat.le_refl 2), ← Nat.cast_add] have h1 := Or.resolve_left (Nat.mod_two_eq_zero_or_one (succ k)) h rw [h1, Nat.add_comm 1, Nat.succ_mul, Nat.one_mul] simp only [Nat.cast_add, Nat.cast_one] @@ -222,18 +222,18 @@ protected def AddMonoidWithOne.binary {R : Type*} [AddMonoid R] [One R] : AddMon #align add_monoid_with_one.binary AddMonoidWithOne.binary theorem one_add_one_eq_two [AddMonoidWithOne α] : 1 + 1 = (2 : α) := by - rw [←Nat.cast_one, ←Nat.cast_add] + rw [← Nat.cast_one, ← Nat.cast_add] apply congrArg decide #align one_add_one_eq_two one_add_one_eq_two theorem two_add_one_eq_three [AddMonoidWithOne α] : 2 + 1 = (3 : α) := by - rw [←one_add_one_eq_two, ←Nat.cast_one, ←Nat.cast_add, ←Nat.cast_add] + rw [← one_add_one_eq_two, ← Nat.cast_one, ← Nat.cast_add, ← Nat.cast_add] apply congrArg decide theorem three_add_one_eq_four [AddMonoidWithOne α] : 3 + 1 = (4 : α) := by - rw [←two_add_one_eq_three, ←one_add_one_eq_two, ←Nat.cast_one, - ←Nat.cast_add, ←Nat.cast_add, ←Nat.cast_add] + rw [← two_add_one_eq_three, ← one_add_one_eq_two, ← Nat.cast_one, + ← Nat.cast_add, ← Nat.cast_add, ← Nat.cast_add] apply congrArg decide diff --git a/Mathlib/Data/Nat/Count.lean b/Mathlib/Data/Nat/Count.lean index d547c042e9ef5..5ae27b3fde4c8 100644 --- a/Mathlib/Data/Nat/Count.lean +++ b/Mathlib/Data/Nat/Count.lean @@ -131,7 +131,7 @@ theorem count_strict_mono {m n : ℕ} (hm : p m) (hmn : m < n) : count p m < cou #align nat.count_strict_mono Nat.count_strict_mono theorem count_injective {m n : ℕ} (hm : p m) (hn : p n) (heq : count p m = count p n) : m = n := by - by_contra' h : m ≠ n + by_contra! h : m ≠ n wlog hmn : m < n · exact this hn hm heq.symm h.symm (h.lt_or_lt.resolve_left hmn) · simpa [heq] using count_strict_mono hm hmn diff --git a/Mathlib/Data/Nat/Digits.lean b/Mathlib/Data/Nat/Digits.lean index 263e583e948c3..03d416d0156de 100644 --- a/Mathlib/Data/Nat/Digits.lean +++ b/Mathlib/Data/Nat/Digits.lean @@ -556,7 +556,7 @@ theorem sub_one_mul_sum_div_pow_eq_sub_sum_digits have w₂' := fun (h : tl ≠ []) ↦ (List.getLast_cons h) ▸ h_ne_zero have ih := ih (w₂' h') w₁' simp only [self_div_pow_eq_ofDigits_drop _ _ h, digits_ofDigits p h tl w₁' w₂', - ←Nat.one_add] at ih + ← Nat.one_add] at ih have := sum_singleton (fun x ↦ ofDigits p <| tl.drop x) tl.length rw [← Ico_succ_singleton, List.drop_length, ofDigits] at this have h₁ : 1 ≤ tl.length := List.length_pos.mpr h' diff --git a/Mathlib/Data/Nat/Factorization/Basic.lean b/Mathlib/Data/Nat/Factorization/Basic.lean index dc024be1680c1..cab8dcfc1d40a 100644 --- a/Mathlib/Data/Nat/Factorization/Basic.lean +++ b/Mathlib/Data/Nat/Factorization/Basic.lean @@ -419,7 +419,7 @@ theorem factorization_le_iff_dvd {d n : ℕ} (hd : d ≠ 0) (hn : n ≠ 0) : set K := n.factorization - d.factorization with hK use K.prod (· ^ ·) rw [← factorization_prod_pow_eq_self hn, ← factorization_prod_pow_eq_self hd, - ←Finsupp.prod_add_index' pow_zero pow_add, hK, add_tsub_cancel_of_le hdn] + ← Finsupp.prod_add_index' pow_zero pow_add, hK, add_tsub_cancel_of_le hdn] · rintro ⟨c, rfl⟩ rw [factorization_mul hd (right_ne_zero_of_mul hn)] simp diff --git a/Mathlib/Data/Nat/Factorization/PrimePow.lean b/Mathlib/Data/Nat/Factorization/PrimePow.lean index 1e0a789d4ef69..b0243819d2132 100644 --- a/Mathlib/Data/Nat/Factorization/PrimePow.lean +++ b/Mathlib/Data/Nat/Factorization/PrimePow.lean @@ -56,7 +56,7 @@ theorem isPrimePow_iff_factorization_eq_single {n : ℕ} : theorem isPrimePow_iff_card_primeFactors_eq_one {n : ℕ} : IsPrimePow n ↔ n.primeFactors.card = 1 := by - simp_rw [isPrimePow_iff_factorization_eq_single, ←Nat.support_factorization, + simp_rw [isPrimePow_iff_factorization_eq_single, ← Nat.support_factorization, Finsupp.card_support_eq_one', exists_prop, pos_iff_ne_zero] #align is_prime_pow_iff_card_support_factorization_eq_one isPrimePow_iff_card_primeFactors_eq_one diff --git a/Mathlib/Data/Nat/Fib/Basic.lean b/Mathlib/Data/Nat/Fib/Basic.lean index 078af88ff9be0..0fe6adaa8a22c 100644 --- a/Mathlib/Data/Nat/Fib/Basic.lean +++ b/Mathlib/Data/Nat/Fib/Basic.lean @@ -285,7 +285,7 @@ theorem gcd_fib_add_self (m n : ℕ) : gcd (fib m) (fib (n + m)) = gcd (fib m) ( theorem gcd_fib_add_mul_self (m n : ℕ) : ∀ k, gcd (fib m) (fib (n + k * m)) = gcd (fib m) (fib n) | 0 => by simp | k + 1 => by - rw [←gcd_fib_add_mul_self m n k, + rw [← gcd_fib_add_mul_self m n k, add_mul, ← add_assoc, one_mul, diff --git a/Mathlib/Data/Nat/Fib/Zeckendorf.lean b/Mathlib/Data/Nat/Fib/Zeckendorf.lean index 19afbb61cfeaf..f88cd9f552728 100644 --- a/Mathlib/Data/Nat/Fib/Zeckendorf.lean +++ b/Mathlib/Data/Nat/Fib/Zeckendorf.lean @@ -55,11 +55,11 @@ lemma IsZeckendorfRep.sum_fib_lt : ∀ {n l}, IsZeckendorfRep l → (∀ a ∈ ( simp only [IsZeckendorfRep, cons_append, chain'_iff_pairwise, pairwise_cons] at hl have : ∀ b, b ∈ head? (l ++ [0]) → b < a - 1 := fun b hb ↦ lt_tsub_iff_right.2 $ hl.1 _ $ mem_of_mem_head? hb - simp only [mem_append, mem_singleton, ←chain'_iff_pairwise, or_imp, forall_and, forall_eq, + simp only [mem_append, mem_singleton, ← chain'_iff_pairwise, or_imp, forall_and, forall_eq, zero_add] at hl simp only [map, List.sum_cons] refine (add_lt_add_left (sum_fib_lt hl.2 this) _).trans_le ?_ - rw [add_comm, ←fib_add_one (hl.1.2.trans_lt' zero_lt_two).ne'] + rw [add_comm, ← fib_add_one (hl.1.2.trans_lt' zero_lt_two).ne'] exact fib_mono (hn _ rfl) end List @@ -100,11 +100,11 @@ lemma greatestFib_ne_zero : greatestFib n ≠ 0 ↔ n ≠ 0 := greatestFib_eq_ze lemma greatestFib_sub_fib_greatestFib_le_greatestFib (hn : n ≠ 0) : greatestFib (n - fib (greatestFib n)) ≤ greatestFib n - 2 := by - rw [←lt_succ_iff, greatestFib_lt, tsub_lt_iff_right n.fib_greatestFib_le, Nat.sub_succ, succ_pred, - ←fib_add_one] + rw [← lt_succ_iff, greatestFib_lt, tsub_lt_iff_right n.fib_greatestFib_le, Nat.sub_succ, + succ_pred, ← fib_add_one] exact n.lt_fib_greatestFib_add_one · simpa - · simpa [←succ_le_iff] using hn.bot_lt + · simpa [← succ_le_iff] using hn.bot_lt private lemma zeckendorf_aux (hm : 0 < m) : m - fib (greatestFib m) < m := tsub_lt_self hm $ fib_pos.2 $ findGreatest_pos.2 ⟨1, zero_lt_one, le_add_self, hm⟩ @@ -150,7 +150,7 @@ lemma zeckendorf_sum_fib : ∀ {l}, IsZeckendorfRep l → zeckendorf (l.map fib) have hl' := hl simp only [IsZeckendorfRep, cons_append, chain'_iff_pairwise, pairwise_cons, mem_append, mem_singleton, or_imp, forall_and, forall_eq, zero_add] at hl - rw [←chain'_iff_pairwise] at hl + rw [← chain'_iff_pairwise] at hl have ha : 0 < a := hl.1.2.trans_lt' zero_lt_two suffices h : greatestFib (fib a + sum (map fib l)) = a · simp only [map, List.sum_cons, add_pos_iff, fib_pos.2 ha, true_or, zeckendorf_of_pos, h, @@ -158,7 +158,7 @@ lemma zeckendorf_sum_fib : ∀ {l}, IsZeckendorfRep l → zeckendorf (l.map fib) simp only [add_comm, add_assoc, greatestFib, findGreatest_eq_iff, ne_eq, ha.ne', not_false_eq_true, le_add_iff_nonneg_left, _root_.zero_le, forall_true_left, not_le, true_and] refine ⟨le_add_of_le_right $ le_fib_add_one _, fun n hn _ ↦ ?_⟩ - rw [add_comm, ←List.sum_cons, ←map_cons] + rw [add_comm, ← List.sum_cons, ← map_cons] exact hl'.sum_fib_lt (by simpa) @[simp] lemma sum_zeckendorf_fib : ∀ n : ℕ, (n.zeckendorf.map fib).sum = n diff --git a/Mathlib/Data/Nat/GCD/Basic.lean b/Mathlib/Data/Nat/GCD/Basic.lean index 844f7e7a69e49..7a7cbdac3ce08 100644 --- a/Mathlib/Data/Nat/GCD/Basic.lean +++ b/Mathlib/Data/Nat/GCD/Basic.lean @@ -135,7 +135,7 @@ theorem lcm_mul_left {m n k : ℕ} : (m * n).lcm (m * k) = m * n.lcm k := by apply dvd_antisymm · exact lcm_dvd (mul_dvd_mul_left m (dvd_lcm_left n k)) (mul_dvd_mul_left m (dvd_lcm_right n k)) · have h : m ∣ lcm (m * n) (m * k) := (dvd_mul_right m n).trans (dvd_lcm_left (m * n) (m * k)) - rw [←dvd_div_iff h, lcm_dvd_iff, dvd_div_iff h, dvd_div_iff h, ←lcm_dvd_iff] + rw [← dvd_div_iff h, lcm_dvd_iff, dvd_div_iff h, dvd_div_iff h, ← lcm_dvd_iff] theorem lcm_mul_right {m n k : ℕ} : (m * n).lcm (k * n) = m.lcm k * n := by rw [mul_comm, mul_comm k n, lcm_mul_left, mul_comm] diff --git a/Mathlib/Data/Nat/MaxPowDiv.lean b/Mathlib/Data/Nat/MaxPowDiv.lean index ef3184fddf3b9..6dd0a360b57c9 100644 --- a/Mathlib/Data/Nat/MaxPowDiv.lean +++ b/Mathlib/Data/Nat/MaxPowDiv.lean @@ -95,7 +95,7 @@ theorem pow_dvd (p n : ℕ) : p ^ (p.maxPowDiv n) ∣ n := by rw [if_pos h] have ⟨c,hc⟩ := pow_dvd p (n / p) rw [go_succ, pow_succ] - nth_rw 2 [←mod_add_div' n p] + nth_rw 2 [← mod_add_div' n p] rw [h.right.right, zero_add] exact ⟨c,by nth_rw 1 [hc]; ac_rfl⟩ · rw [if_neg h] diff --git a/Mathlib/Data/Nat/Order/Lemmas.lean b/Mathlib/Data/Nat/Order/Lemmas.lean index 347d22795547b..53533ef5aaddf 100644 --- a/Mathlib/Data/Nat/Order/Lemmas.lean +++ b/Mathlib/Data/Nat/Order/Lemmas.lean @@ -223,11 +223,6 @@ theorem le_of_lt_add_of_dvd (h : a < b + n) : n ∣ a → n ∣ b → a ≤ b := exact mul_le_mul_left' (lt_succ_iff.1 <| lt_of_mul_lt_mul_left h bot_le) _ #align nat.le_of_lt_add_of_dvd Nat.le_of_lt_add_of_dvd -@[simp] -theorem mod_div_self (m n : ℕ) : m % n / n = 0 := by - cases n - · exact (m % 0).div_zero - · case succ n => exact Nat.div_eq_of_lt (m.mod_lt n.succ_pos) #align nat.mod_div_self Nat.mod_div_self /-- `n` is not divisible by `a` iff it is between `a * k` and `a * (k + 1)` for some `k`. -/ diff --git a/Mathlib/Data/Nat/Parity.lean b/Mathlib/Data/Nat/Parity.lean index e8329207d53b3..1af9d26002a70 100644 --- a/Mathlib/Data/Nat/Parity.lean +++ b/Mathlib/Data/Nat/Parity.lean @@ -305,7 +305,7 @@ theorem iterate_bit0 (hf : Involutive f) (n : ℕ) : f^[bit0 n] = id := by #align function.involutive.iterate_bit0 Function.Involutive.iterate_bit0 theorem iterate_bit1 (hf : Involutive f) (n : ℕ) : f^[bit1 n] = f := by - rw [bit1, ←succ_eq_add_one, iterate_succ, hf.iterate_bit0, comp.left_id] + rw [bit1, ← succ_eq_add_one, iterate_succ, hf.iterate_bit0, comp.left_id] #align function.involutive.iterate_bit1 Function.Involutive.iterate_bit1 end diff --git a/Mathlib/Data/Nat/PartENat.lean b/Mathlib/Data/Nat/PartENat.lean index 3d959a474bd46..05bb8d9ce9ae9 100644 --- a/Mathlib/Data/Nat/PartENat.lean +++ b/Mathlib/Data/Nat/PartENat.lean @@ -561,7 +561,7 @@ theorem add_eq_top_iff {a b : PartENat} : a + b = ⊤ ↔ a = ⊤ ∨ b = ⊤ := refine PartENat.casesOn a ?_ ?_ <;> refine PartENat.casesOn b ?_ ?_ <;> simp [top_add, add_top] - simp only [←Nat.cast_add, PartENat.natCast_ne_top, forall_const, not_false_eq_true] + simp only [← Nat.cast_add, PartENat.natCast_ne_top, forall_const, not_false_eq_true] #align part_enat.add_eq_top_iff PartENat.add_eq_top_iff protected theorem add_right_cancel_iff {a b c : PartENat} (hc : c ≠ ⊤) : a + c = b + c ↔ a = b := by @@ -569,7 +569,7 @@ protected theorem add_right_cancel_iff {a b c : PartENat} (hc : c ≠ ⊤) : a + refine PartENat.casesOn a ?_ ?_ <;> refine PartENat.casesOn b ?_ ?_ <;> simp [add_eq_top_iff, natCast_ne_top, @eq_comm _ (⊤ : PartENat), top_add] - simp only [←Nat.cast_add, add_left_cancel_iff, PartENat.natCast_inj, add_comm, forall_const] + simp only [← Nat.cast_add, add_left_cancel_iff, PartENat.natCast_inj, add_comm, forall_const] #align part_enat.add_right_cancel_iff PartENat.add_right_cancel_iff protected theorem add_left_cancel_iff {a b c : PartENat} (ha : a ≠ ⊤) : a + b = a + c ↔ b = c := by diff --git a/Mathlib/Data/Nat/Prime.lean b/Mathlib/Data/Nat/Prime.lean index 2347dcc11e9f6..bc3d5c3f5537d 100644 --- a/Mathlib/Data/Nat/Prime.lean +++ b/Mathlib/Data/Nat/Prime.lean @@ -174,7 +174,7 @@ theorem prime_three : Prime 3 := by decide theorem Prime.five_le_of_ne_two_of_ne_three {p : ℕ} (hp : p.Prime) (h_two : p ≠ 2) (h_three : p ≠ 3) : 5 ≤ p := by - by_contra' h + by_contra! h revert h_two h_three hp -- Porting note: was `decide!` match p with @@ -430,7 +430,7 @@ theorem minFac_eq_one_iff {n : ℕ} : minFac n = 1 ↔ n = 1 := by theorem minFac_eq_two_iff (n : ℕ) : minFac n = 2 ↔ 2 ∣ n := by constructor · intro h - rw [←h] + rw [← h] exact minFac_dvd n · intro h have ub := minFac_le_of_dvd (le_refl 2) h diff --git a/Mathlib/Data/Nat/PrimeFin.lean b/Mathlib/Data/Nat/PrimeFin.lean index 09af33f845e21..1a2efccc29f54 100644 --- a/Mathlib/Data/Nat/PrimeFin.lean +++ b/Mathlib/Data/Nat/PrimeFin.lean @@ -35,7 +35,7 @@ instance Primes.countable : Countable Primes := ⟨⟨coeNat.coe, coe_nat_inject @[simp] lemma toFinset_factors (n : ℕ) : n.factors.toFinset = n.primeFactors := rfl @[simp] lemma mem_primeFactors : p ∈ n.primeFactors ↔ p.Prime ∧ p ∣ n ∧ n ≠ 0 := by - simp_rw [←toFinset_factors, List.mem_toFinset, mem_factors'] + simp_rw [← toFinset_factors, List.mem_toFinset, mem_factors'] lemma mem_primeFactors_of_ne_zero (hn : n ≠ 0) : p ∈ n.primeFactors ↔ p.Prime ∧ p ∣ n := by simp [hn] @@ -86,7 +86,7 @@ lemma primeFactors_gcd (ha : a ≠ 0) (hb : b ≠ 0) : @[simp] lemma disjoint_primeFactors (ha : a ≠ 0) (hb : b ≠ 0) : Disjoint a.primeFactors b.primeFactors ↔ Coprime a b := by - simp [disjoint_iff_inter_eq_empty, coprime_iff_gcd_eq_one, ←primeFactors_gcd, gcd_ne_zero_left, + simp [disjoint_iff_inter_eq_empty, coprime_iff_gcd_eq_one, ← primeFactors_gcd, gcd_ne_zero_left, ha, hb] protected lemma Coprime.disjoint_primeFactors (hab : Coprime a b) : diff --git a/Mathlib/Data/Nat/Size.lean b/Mathlib/Data/Nat/Size.lean index 3a9c95879822e..0feaee57900ea 100644 --- a/Mathlib/Data/Nat/Size.lean +++ b/Mathlib/Data/Nat/Size.lean @@ -32,10 +32,7 @@ theorem shiftLeft'_tt_eq_mul_pow (m) : ∀ n, shiftLeft' true m n + 1 = (m + 1) end #align nat.one_shiftl Nat.one_shiftLeft - -theorem zero_shiftLeft (n) : 0 <<< n = 0 := by simp #align nat.zero_shiftl Nat.zero_shiftLeft - #align nat.shiftr_eq_div_pow Nat.shiftRight_eq_div_pow theorem shiftLeft'_ne_zero_left (b) {m} (h : m ≠ 0) (n) : shiftLeft' b m n ≠ 0 := by @@ -118,7 +115,7 @@ theorem lt_size_self (n : ℕ) : n < 2 ^ size n := by by_cases h : bit b n = 0 · apply this h rw [size_bit h, shiftLeft_succ, shiftLeft_eq, one_mul, ← bit0_val] - exact bit_lt_bit0 _ (by simpa [shiftRight_eq_div_pow] using IH) + exact bit_lt_bit0 _ (by simpa [shiftLeft_eq, shiftRight_eq_div_pow] using IH) #align nat.lt_size_self Nat.lt_size_self theorem size_le {m n : ℕ} : size m ≤ n ↔ m < 2 ^ n := diff --git a/Mathlib/Data/Nat/Sqrt.lean b/Mathlib/Data/Nat/Sqrt.lean index cbd3e6973d366..a56d4127ff713 100644 --- a/Mathlib/Data/Nat/Sqrt.lean +++ b/Mathlib/Data/Nat/Sqrt.lean @@ -180,7 +180,7 @@ theorem exists_mul_self' (x : ℕ) : (∃ n, n ^ 2 = x) ↔ sqrt x ^ 2 = x := by /-- `IsSquare` can be decided on `ℕ` by checking against the square root. -/ instance : DecidablePred (IsSquare : ℕ → Prop) := fun m => decidable_of_iff' (Nat.sqrt m * Nat.sqrt m = m) <| by - simp_rw [←Nat.exists_mul_self m, IsSquare, eq_comm] + simp_rw [← Nat.exists_mul_self m, IsSquare, eq_comm] theorem sqrt_mul_sqrt_lt_succ (n : ℕ) : sqrt n * sqrt n < n + 1 := lt_succ_iff.mpr (sqrt_le _) diff --git a/Mathlib/Data/Nat/Squarefree.lean b/Mathlib/Data/Nat/Squarefree.lean index cb49fb02d565c..24f6b53a0e8a3 100644 --- a/Mathlib/Data/Nat/Squarefree.lean +++ b/Mathlib/Data/Nat/Squarefree.lean @@ -410,7 +410,7 @@ lemma primeFactors_prod (hs : ∀ p ∈ s, p.Prime) : primeFactors (∏ p in s, rw [mem_primeFactors_of_ne_zero hn, and_congr_right (fun hp ↦ hp.prime.dvd_finset_prod_iff _)] refine' ⟨_, fun hp ↦ ⟨hs _ hp, _, hp, dvd_rfl⟩⟩ rintro ⟨hp, q, hq, hpq⟩ - rwa [←((hs _ hq).dvd_iff_eq hp.ne_one).1 hpq] + rwa [← ((hs _ hq).dvd_iff_eq hp.ne_one).1 hpq] lemma primeFactors_div_gcd (hm : Squarefree m) (hn : n ≠ 0) : primeFactors (m / m.gcd n) = primeFactors m \ primeFactors n := by diff --git a/Mathlib/Data/Num/Lemmas.lean b/Mathlib/Data/Num/Lemmas.lean index dc9ced5fbb7cf..ef4f5a2404d02 100644 --- a/Mathlib/Data/Num/Lemmas.lean +++ b/Mathlib/Data/Num/Lemmas.lean @@ -420,7 +420,7 @@ instance commSemiring : CommSemiring Num := by instance orderedCancelAddCommMonoid : OrderedCancelAddCommMonoid Num where le := (· ≤ ·) lt := (· < ·) - lt_iff_le_not_le a b := by simp only [←lt_to_nat, ←le_to_nat, lt_iff_le_not_le] + lt_iff_le_not_le a b := by simp only [← lt_to_nat, ← le_to_nat, lt_iff_le_not_le] le_refl := by transfer le_trans a b c := by transfer_rw; apply le_trans le_antisymm a b := by transfer_rw; apply le_antisymm @@ -897,8 +897,8 @@ theorem castNum_eq_bitwise {f : Num → Num → Num} {g : Bool → Bool → Bool intros b _; cases b <;> rfl] rw [Nat.bitwise_bit gff] any_goals rw [Nat.bitwise_zero, p11]; cases g true true <;> rfl - any_goals rw [Nat.bitwise_zero_left, ←Bool.cond_eq_ite, this, ← bit_to_nat, p1b] - any_goals rw [Nat.bitwise_zero_right, ←Bool.cond_eq_ite, this, ← bit_to_nat, pb1] + any_goals rw [Nat.bitwise_zero_left, ← Bool.cond_eq_ite, this, ← bit_to_nat, p1b] + any_goals rw [Nat.bitwise_zero_right, ← Bool.cond_eq_ite, this, ← bit_to_nat, pb1] all_goals rw [← show ∀ n : PosNum, ↑(p m n) = Nat.bitwise g ↑m ↑n from IH] rw [← bit_to_nat, pbb] @@ -929,7 +929,7 @@ theorem castNum_xor : ∀ m n : Num, ↑(m ^^^ n) = (↑m ^^^ ↑n : ℕ) := by @[simp, norm_cast] theorem castNum_shiftLeft (m : Num) (n : Nat) : ↑(m <<< n) = (m : ℕ) <<< (n : ℕ) := by - cases m <;> dsimp only [←shiftl_eq_shiftLeft, shiftl] + cases m <;> dsimp only [← shiftl_eq_shiftLeft, shiftl] · symm apply Nat.zero_shiftLeft simp only [cast_pos] @@ -943,12 +943,12 @@ theorem castNum_shiftLeft (m : Num) (n : Nat) : ↑(m <<< n) = (m : ℕ) <<< (n @[simp, norm_cast] theorem castNum_shiftRight (m : Num) (n : Nat) : ↑(m >>> n) = (m : ℕ) >>> (n : ℕ) := by - cases' m with m <;> dsimp only [←shiftr_eq_shiftRight, shiftr]; + cases' m with m <;> dsimp only [← shiftr_eq_shiftRight, shiftr]; · symm apply Nat.zero_shiftRight induction' n with n IH generalizing m · cases m <;> rfl - cases' m with m m <;> dsimp only [PosNum.shiftr, ←PosNum.shiftr_eq_shiftRight] + cases' m with m m <;> dsimp only [PosNum.shiftr, ← PosNum.shiftr_eq_shiftRight] · rw [Nat.shiftRight_eq_div_pow] symm apply Nat.div_eq_of_lt @@ -970,28 +970,19 @@ theorem castNum_shiftRight (m : Num) (n : Nat) : ↑(m >>> n) = (m : ℕ) >>> (n @[simp] theorem castNum_testBit (m n) : testBit m n = Nat.testBit m n := by -- Porting note: `unfold` → `dsimp only` - cases m <;> dsimp only [testBit, Nat.testBit] + cases m <;> dsimp only [testBit] case zero => - change false = Nat.bodd (0 >>> n) - rw [Nat.zero_shiftRight] - rfl + rw [show (Num.zero : Nat) = 0 from rfl, Nat.zero_testBit] case pos m => - induction' n with n IH generalizing m <;> cases' m with m m <;> dsimp only [PosNum.testBit] + rw [cast_pos] + induction' n with n IH generalizing m <;> cases' m with m m + <;> dsimp only [PosNum.testBit, Nat.zero_eq] · rfl - · exact (Nat.bodd_bit _ _).symm - · exact (Nat.bodd_bit _ _).symm - · change false = Nat.bodd (1 >>> (n + 1)) - rw [add_comm, Nat.shiftRight_add] - change false = Nat.bodd (0 >>> n) - rw [Nat.zero_shiftRight]; rfl - · change PosNum.testBit m n = Nat.bodd ((Nat.bit true m) >>> (n + 1)) - rw [add_comm, Nat.shiftRight_add] - simp only [Nat.shiftRight_succ, Nat.shiftRight_zero, ← Nat.div2_val, Nat.div2_bit] - apply IH - · change PosNum.testBit m n = Nat.bodd ((Nat.bit false m) >>> (n + 1)) - rw [add_comm, Nat.shiftRight_add] - simp only [Nat.shiftRight_succ, Nat.shiftRight_zero, ← Nat.div2_val, Nat.div2_bit] - apply IH + · rw [PosNum.cast_bit1, ← Nat.bit_true, Nat.testBit_zero] + · rw [PosNum.cast_bit0, ← Nat.bit_false, Nat.testBit_zero] + · rw [PosNum.cast_one', ← bit1_zero, ← Nat.bit_true, Nat.testBit_succ, Nat.zero_testBit] + · rw [PosNum.cast_bit1, ← Nat.bit_true, Nat.testBit_succ, IH] + · rw [PosNum.cast_bit0, ← Nat.bit_false, Nat.testBit_succ, IH] #align num.test_bit_to_nat Num.castNum_testBit end Num diff --git a/Mathlib/Data/PEquiv.lean b/Mathlib/Data/PEquiv.lean index b2958bfe7fb52..df73a60fe81ce 100644 --- a/Mathlib/Data/PEquiv.lean +++ b/Mathlib/Data/PEquiv.lean @@ -135,8 +135,11 @@ theorem symm_refl : (PEquiv.refl α).symm = PEquiv.refl α := theorem symm_symm (f : α ≃. β) : f.symm.symm = f := by cases f; rfl #align pequiv.symm_symm PEquiv.symm_symm +theorem symm_bijective : Function.Bijective (PEquiv.symm : (α ≃. β) → β ≃. α) := + Function.bijective_iff_has_inverse.mpr ⟨_, symm_symm, symm_symm⟩ + theorem symm_injective : Function.Injective (@PEquiv.symm α β) := - LeftInverse.injective symm_symm + symm_bijective.injective #align pequiv.symm_injective PEquiv.symm_injective theorem trans_assoc (f : α ≃. β) (g : β ≃. γ) (h : γ ≃. δ) : diff --git a/Mathlib/Data/PNat/Interval.lean b/Mathlib/Data/PNat/Interval.lean index 8a0520c6661b8..dbd79d4263262 100644 --- a/Mathlib/Data/PNat/Interval.lean +++ b/Mathlib/Data/PNat/Interval.lean @@ -102,7 +102,7 @@ theorem card_Ioo : (Ioo a b).card = b - a - 1 := by @[simp] theorem card_uIcc : (uIcc a b).card = (b - a : ℤ).natAbs + 1 := by - rw [←Nat.card_uIcc, ←map_subtype_embedding_uIcc, card_map] + rw [← Nat.card_uIcc, ← map_subtype_embedding_uIcc, card_map] #align pnat.card_uIcc PNat.card_uIcc -- porting note: `simpNF` says `simp` can prove this @@ -127,7 +127,7 @@ theorem card_fintype_Ioo : Fintype.card (Set.Ioo a b) = b - a - 1 := by -- porting note: `simpNF` says `simp` can prove this theorem card_fintype_uIcc : Fintype.card (Set.uIcc a b) = (b - a : ℤ).natAbs + 1 := by - rw [←card_uIcc, Fintype.card_ofFinset] + rw [← card_uIcc, Fintype.card_ofFinset] #align pnat.card_fintype_uIcc PNat.card_fintype_uIcc end PNat diff --git a/Mathlib/Data/Polynomial/Degree/Definitions.lean b/Mathlib/Data/Polynomial/Degree/Definitions.lean index 931e80fb15047..06eadf6592a09 100644 --- a/Mathlib/Data/Polynomial/Degree/Definitions.lean +++ b/Mathlib/Data/Polynomial/Degree/Definitions.lean @@ -635,7 +635,7 @@ theorem degree_le_zero_iff : degree p ≤ 0 ↔ p = C (coeff p 0) := #align polynomial.degree_le_zero_iff Polynomial.degree_le_zero_iff theorem degree_add_le (p q : R[X]) : degree (p + q) ≤ max (degree p) (degree q) := by - simpa only [degree, ←support_toFinsupp, toFinsupp_add] + simpa only [degree, ← support_toFinsupp, toFinsupp_add] using AddMonoidAlgebra.sup_support_add_le _ _ _ #align polynomial.degree_add_le Polynomial.degree_add_le @@ -804,7 +804,7 @@ theorem degree_sum_le (s : Finset ι) (f : ι → R[X]) : #align polynomial.degree_sum_le Polynomial.degree_sum_le theorem degree_mul_le (p q : R[X]) : degree (p * q) ≤ degree p + degree q := by - simpa only [degree, ←support_toFinsupp, toFinsupp_mul] + simpa only [degree, ← support_toFinsupp, toFinsupp_mul] using AddMonoidAlgebra.sup_support_mul_le (WithBot.coe_add _ _).le _ _ #align polynomial.degree_mul_le Polynomial.degree_mul_le @@ -957,7 +957,7 @@ theorem coeff_mul_degree_add_degree (p q : R[X]) : rw [coeff_eq_zero_of_degree_lt (lt_of_le_of_lt degree_le_natDegree (WithBot.coe_lt_coe.2 this)), mul_zero] - · by_contra' H' + · by_contra! H' exact ne_of_lt (Nat.lt_of_lt_of_le (Nat.add_lt_add_right H j) (Nat.add_le_add_left H' _)) h₁ diff --git a/Mathlib/Data/Polynomial/Degree/TrailingDegree.lean b/Mathlib/Data/Polynomial/Degree/TrailingDegree.lean index 507f2a4792a0c..2afe0949fffd7 100644 --- a/Mathlib/Data/Polynomial/Degree/TrailingDegree.lean +++ b/Mathlib/Data/Polynomial/Degree/TrailingDegree.lean @@ -177,7 +177,7 @@ theorem trailingDegree_ne_of_natTrailingDegree_ne {n : ℕ} : p.natTrailingDegree ≠ n → trailingDegree p ≠ n := by -- Porting note: Needed to account for different coercion behaviour & add the lemma below have : Nat.cast n = WithTop.some n := rfl - exact mt fun h => by rw [natTrailingDegree, h, this, ←WithTop.some_eq_coe, Option.getD_some] + exact mt fun h => by rw [natTrailingDegree, h, this, ← WithTop.some_eq_coe, Option.getD_some] #align polynomial.trailing_degree_ne_of_nat_trailing_degree_ne Polynomial.trailingDegree_ne_of_natTrailingDegree_ne theorem natTrailingDegree_le_of_trailingDegree_le {n : ℕ} {hp : p ≠ 0} diff --git a/Mathlib/Data/Polynomial/DenomsClearable.lean b/Mathlib/Data/Polynomial/DenomsClearable.lean index 8ebf930d50de3..e965da870d18f 100644 --- a/Mathlib/Data/Polynomial/DenomsClearable.lean +++ b/Mathlib/Data/Polynomial/DenomsClearable.lean @@ -100,7 +100,7 @@ theorem one_le_pow_mul_abs_eval_div {K : Type*} [LinearOrderedField K] {f : ℤ[ obtain Fa := _root_.congr_arg abs hF rw [eq_one_div_of_mul_eq_one_left bu, eq_intCast, eq_intCast, abs_mul] at Fa rw [abs_of_pos (pow_pos (Int.cast_pos.mpr b0) _ : 0 < (b : K) ^ _), one_div, eq_intCast] at Fa - rw [div_eq_mul_inv, ←Fa, ← Int.cast_abs, ← Int.cast_one, Int.cast_le] + rw [div_eq_mul_inv, ← Fa, ← Int.cast_abs, ← Int.cast_one, Int.cast_le] refine' Int.le_of_lt_add_one ((lt_add_iff_pos_left 1).mpr (abs_pos.mpr fun F0 => fab _)) rw [eq_one_div_of_mul_eq_one_left bu, F0, one_div, eq_intCast, Int.cast_zero, zero_eq_mul] at hF cases' hF with hF hF diff --git a/Mathlib/Data/Polynomial/Derivation.lean b/Mathlib/Data/Polynomial/Derivation.lean index 100ef936804c5..1a589a6818b96 100644 --- a/Mathlib/Data/Polynomial/Derivation.lean +++ b/Mathlib/Data/Polynomial/Derivation.lean @@ -131,7 +131,7 @@ def compAEval : Derivation R R[X] <| AEval R M a where -/ theorem compAEval_eq (d : Derivation R A M) (f : R[X]) : d.compAEval a f = derivative f • (AEval.of R M a (d a)) := by - rw [←mkDerivation_apply] + rw [← mkDerivation_apply] congr apply derivation_ext simp diff --git a/Mathlib/Data/Polynomial/Derivative.lean b/Mathlib/Data/Polynomial/Derivative.lean index e3c921c66158a..6fb69830afa27 100644 --- a/Mathlib/Data/Polynomial/Derivative.lean +++ b/Mathlib/Data/Polynomial/Derivative.lean @@ -273,7 +273,7 @@ theorem natDegree_eq_zero_of_derivative_eq_zero [NoZeroSMulDivisors ℕ R] {f : rcases eq_or_ne f 0 with (rfl | hf) · exact natDegree_zero rw [natDegree_eq_zero_iff_degree_le_zero] - by_contra' f_nat_degree_pos + by_contra! f_nat_degree_pos rw [← natDegree_pos_iff_degree_pos] at f_nat_degree_pos let m := f.natDegree - 1 have hm : m + 1 = f.natDegree := tsub_add_cancel_of_le f_nat_degree_pos diff --git a/Mathlib/Data/Polynomial/Eval.lean b/Mathlib/Data/Polynomial/Eval.lean index 5ec2eba436d0e..f5f467517ddc2 100644 --- a/Mathlib/Data/Polynomial/Eval.lean +++ b/Mathlib/Data/Polynomial/Eval.lean @@ -156,7 +156,7 @@ theorem eval₂_sum (p : T[X]) (g : ℕ → T → R[X]) (x : S) : map_add' := fun p q => eval₂_add _ _ } have A : ∀ y, eval₂ f x y = T y := fun y => rfl simp only [A] - rw [sum, T.map_sum, sum] + rw [sum, map_sum, sum] #align polynomial.eval₂_sum Polynomial.eval₂_sum theorem eval₂_list_sum (l : List R[X]) (x : S) : eval₂ f x l.sum = (l.map (eval₂ f x)).sum := diff --git a/Mathlib/Data/Polynomial/FieldDivision.lean b/Mathlib/Data/Polynomial/FieldDivision.lean index 6c6616c51ff4c..9f8dde57e030c 100644 --- a/Mathlib/Data/Polynomial/FieldDivision.lean +++ b/Mathlib/Data/Polynomial/FieldDivision.lean @@ -402,7 +402,7 @@ theorem exists_root_of_degree_eq_one (h : degree p = 1) : ∃ x, IsRoot p x := ⟨-(p.coeff 0 / p.coeff 1), by have : p.coeff 1 ≠ 0 := by have h' := natDegree_eq_of_degree_eq_some h - change natDegree p = 1 at h'; rw [←h'] + change natDegree p = 1 at h'; rw [← h'] exact mt leadingCoeff_eq_zero.1 fun h0 => by simp [h0] at h; contradiction conv in p => rw [eq_X_add_C_of_degree_le_one (show degree p ≤ 1 by rw [h])] simp [IsRoot, mul_div_cancel' _ this]⟩ @@ -533,7 +533,7 @@ theorem X_sub_C_dvd_derivative_of_X_sub_C_dvd_divByMonic {K : Type*} [Field K] ( (hf : (X - C a) ∣ f /ₘ (X - C a)) : X - C a ∣ derivative f := by have key := divByMonic_add_X_sub_C_mul_derivate_divByMonic_eq_derivative f a have ⟨u,hu⟩ := hf - rw [←key, hu, ←mul_add (X - C a) u _] + rw [← key, hu, ← mul_add (X - C a) u _] use (u + derivative ((X - C a) * u)) /-- If `f` is a polynomial over a field, and `a : K` satisfies `f' a ≠ 0`, diff --git a/Mathlib/Data/Polynomial/HasseDeriv.lean b/Mathlib/Data/Polynomial/HasseDeriv.lean index d909e1b11a6b0..db045964e9776 100644 --- a/Mathlib/Data/Polynomial/HasseDeriv.lean +++ b/Mathlib/Data/Polynomial/HasseDeriv.lean @@ -158,7 +158,7 @@ theorem factorial_smul_hasseDeriv : ⇑(k ! • @hasseDeriv R _ k) = (@derivativ rw [mul_comm (k+1) _, mul_assoc, mul_assoc] congr 1 have : n + k + 1 = n + (k + 1) := by apply add_assoc - rw [←choose_symm_of_eq_add this, choose_succ_right_eq, mul_comm] + rw [← choose_symm_of_eq_add this, choose_succ_right_eq, mul_comm] congr rw [add_assoc, add_tsub_cancel_left] #align polynomial.factorial_smul_hasse_deriv Polynomial.factorial_smul_hasseDeriv @@ -170,7 +170,7 @@ theorem hasseDeriv_comp (k l : ℕ) : mul_one, monomial_eq_zero_iff, sum_monomial_index, mul_zero, ← tsub_add_eq_tsub_tsub, add_comm l k] rw_mod_cast [nsmul_eq_mul] - rw [←Nat.cast_mul] + rw [← Nat.cast_mul] congr 2 by_cases hikl : i < k + l · rw [choose_eq_zero_of_lt hikl, mul_zero] @@ -263,7 +263,7 @@ theorem hasseDeriv_mul (f g : R[X]) : rw [Finset.sum_congr rfl aux] rw [← map_sum, ← Finset.sum_mul] congr - rw_mod_cast [←Nat.add_choose_eq] + rw_mod_cast [← Nat.add_choose_eq] #align polynomial.hasse_deriv_mul Polynomial.hasseDeriv_mul end diff --git a/Mathlib/Data/Polynomial/Lifts.lean b/Mathlib/Data/Polynomial/Lifts.lean index 6b3aef253ab5f..3f4e2823d475d 100644 --- a/Mathlib/Data/Polynomial/Lifts.lean +++ b/Mathlib/Data/Polynomial/Lifts.lean @@ -194,11 +194,11 @@ theorem mem_lifts_and_degree_eq {p : S[X]} (hlifts : p ∈ lifts f) : use erase + lead constructor · simp only [hlead, herase, Polynomial.map_add] - rw [←eraseLead, ←leadingCoeff] + rw [← eraseLead, ← leadingCoeff] rw [eraseLead_add_monomial_natDegree_leadingCoeff p] - rw [degree_eq_natDegree pzero, ←deg_lead] + rw [degree_eq_natDegree pzero, ← deg_lead] apply degree_add_eq_right_of_degree_lt - rw [herase.2, deg_lead, ←degree_eq_natDegree pzero] + rw [herase.2, deg_lead, ← degree_eq_natDegree pzero] exact degree_erase_lt pzero #align polynomial.mem_lifts_and_degree_eq Polynomial.mem_lifts_and_degree_eq @@ -223,7 +223,7 @@ theorem lifts_and_degree_eq_and_monic [Nontrivial S] {p : S[X]} (hlifts : p ∈ obtain ⟨q, hq⟩ := mem_lifts_and_degree_eq (erase_mem_lifts p.natDegree hlifts) have p_neq_0 : p ≠ 0 := by intro hp; apply h0; rw [hp]; simp only [natDegree_zero, erase_zero] have hdeg : q.degree < (X ^ p.natDegree).degree := by - rw [@degree_X_pow R, hq.2, ←degree_eq_natDegree p_neq_0] + rw [@degree_X_pow R, hq.2, ← degree_eq_natDegree p_neq_0] exact degree_erase_lt p_neq_0 refine' ⟨q + X ^ p.natDegree, _, _, (monic_X_pow _).add_of_right hdeg⟩ · rw [Polynomial.map_add, hq.1, Polynomial.map_pow, map_X, H] diff --git a/Mathlib/Data/Polynomial/Module.lean b/Mathlib/Data/Polynomial/Module.lean index 117636f6122df..c1c4bd6014e59 100644 --- a/Mathlib/Data/Polynomial/Module.lean +++ b/Mathlib/Data/Polynomial/Module.lean @@ -70,7 +70,7 @@ lemma of_aeval_smul (f : R[X]) (m : M) : of R M a (aeval a f • m) = f • of R (of R M a).symm (f • m) = aeval a f • (of R M a).symm m := rfl lemma X_smul_of (m : M) : (X : R[X]) • (of R M a m) = of R M a (a • m) := by - rw [←of_aeval_smul, aeval_X] + rw [← of_aeval_smul, aeval_X] lemma of_symm_X_smul (m : AEval R M a) : (of R M a).symm ((X : R[X]) • m) = a • (of R M a).symm m := by @@ -306,7 +306,7 @@ noncomputable def equivPolynomialSelf : PolynomialModule R R ≃ₗ[R[X]] R[X] : · dsimp at H exfalso apply hpq2 - rw [←hpq1, H] + rw [← hpq1, H] simp only [add_le_iff_nonpos_left, nonpos_iff_eq_zero, add_tsub_cancel_right] · rfl · intro H diff --git a/Mathlib/Data/Polynomial/RingDivision.lean b/Mathlib/Data/Polynomial/RingDivision.lean index b4283a76d9069..ae9978ac7c2e2 100644 --- a/Mathlib/Data/Polynomial/RingDivision.lean +++ b/Mathlib/Data/Polynomial/RingDivision.lean @@ -139,8 +139,8 @@ instance : NoZeroDivisors R[X] where rw [← leadingCoeff_zero, ← leadingCoeff_mul, h] theorem natDegree_mul (hp : p ≠ 0) (hq : q ≠ 0) : (p*q).natDegree = p.natDegree + q.natDegree := by - rw [← Nat.cast_inj (R := WithBot ℕ), ←degree_eq_natDegree (mul_ne_zero hp hq), - Nat.cast_add, ←degree_eq_natDegree hp, ← degree_eq_natDegree hq, degree_mul] + rw [← Nat.cast_inj (R := WithBot ℕ), ← degree_eq_natDegree (mul_ne_zero hp hq), + Nat.cast_add, ← degree_eq_natDegree hp, ← degree_eq_natDegree hq, degree_mul] #align polynomial.nat_degree_mul Polynomial.natDegree_mul theorem trailingDegree_mul : (p * q).trailingDegree = p.trailingDegree + q.trailingDegree := by @@ -860,6 +860,29 @@ theorem mem_nthRootsFinset {n : ℕ} (h : 0 < n) {x : R} : theorem nthRootsFinset_zero : nthRootsFinset 0 R = ∅ := by classical simp [nthRootsFinset_def] #align polynomial.nth_roots_finset_zero Polynomial.nthRootsFinset_zero +theorem mul_mem_nthRootsFinset + {η₁ η₂ : R} (hη₁ : η₁ ∈ nthRootsFinset n R) (hη₂ : η₂ ∈ nthRootsFinset n R) : + η₁ * η₂ ∈ nthRootsFinset n R := by + cases n with + | zero => + simp only [Nat.zero_eq, nthRootsFinset_zero, not_mem_empty] at hη₁ + | succ n => + rw [mem_nthRootsFinset n.succ_pos] at hη₁ hη₂ ⊢ + rw [mul_pow, hη₁, hη₂, one_mul] + +theorem ne_zero_of_mem_nthRootsFinset {η : R} (hη : η ∈ nthRootsFinset n R) : η ≠ 0 := by + nontriviality R + rintro rfl + cases n with + | zero => + simp only [Nat.zero_eq, nthRootsFinset_zero, not_mem_empty] at hη + | succ n => + rw [mem_nthRootsFinset n.succ_pos, zero_pow n.succ_pos] at hη + exact zero_ne_one hη + +theorem one_mem_nthRootsFinset (hn : 0 < n) : 1 ∈ nthRootsFinset n R := by + rw [mem_nthRootsFinset hn, one_pow] + end NthRoots theorem Monic.comp (hp : p.Monic) (hq : q.Monic) (h : q.natDegree ≠ 0) : (p.comp q).Monic := by @@ -1172,7 +1195,7 @@ theorem isCoprime_X_sub_C_of_isUnit_sub {R} [CommRing R] {a b : R} (h : IsUnit ( IsCoprime (X - C a) (X - C b) := ⟨-C h.unit⁻¹.val, C h.unit⁻¹.val, by rw [neg_mul_comm, ← left_distrib, neg_add_eq_sub, sub_sub_sub_cancel_left, ← C_sub, ← C_mul] - rw [←C_1] + rw [← C_1] congr exact h.val_inv_mul⟩ set_option linter.uppercaseLean3 false in diff --git a/Mathlib/Data/QPF/Multivariate/Constructions/Cofix.lean b/Mathlib/Data/QPF/Multivariate/Constructions/Cofix.lean index 3862cc4a8fb65..695ef5d7fedcd 100644 --- a/Mathlib/Data/QPF/Multivariate/Constructions/Cofix.lean +++ b/Mathlib/Data/QPF/Multivariate/Constructions/Cofix.lean @@ -442,7 +442,7 @@ theorem Cofix.abs_repr {α} (x : Cofix F α) : Quot.mk _ (Cofix.repr x) = x := b congr rfl rw [Cofix.dest] - rw [MvFunctor.map_map, MvFunctor.map_map, ←appendFun_comp_id, ←appendFun_comp_id] + rw [MvFunctor.map_map, MvFunctor.map_map, ← appendFun_comp_id, ← appendFun_comp_id] apply liftR_map_last intros rfl diff --git a/Mathlib/Data/Rat/Defs.lean b/Mathlib/Data/Rat/Defs.lean index c382e424889b0..654b9dbe6abfe 100644 --- a/Mathlib/Data/Rat/Defs.lean +++ b/Mathlib/Data/Rat/Defs.lean @@ -87,7 +87,7 @@ lemma num_ne_zero {q : ℚ} : q.num ≠ 0 ↔ q ≠ 0 := num_eq_zero.not @[simp] theorem divInt_eq_zero {a b : ℤ} (b0 : b ≠ 0) : a /. b = 0 ↔ a = 0 := by - rw [←zero_divInt b, divInt_eq_iff b0 b0, zero_mul, mul_eq_zero, or_iff_left b0] + rw [← zero_divInt b, divInt_eq_iff b0 b0, zero_mul, mul_eq_zero, or_iff_left b0] #align rat.mk_eq_zero Rat.divInt_eq_zero theorem divInt_ne_zero {a b : ℤ} (b0 : b ≠ 0) : a /. b ≠ 0 ↔ a ≠ 0 := @@ -557,7 +557,7 @@ theorem mkRat_eq_div {n : ℤ} {d : ℕ} : mkRat n d = n / d := by unfold Rat.inv have h₁ : 0 < d := Nat.pos_iff_ne_zero.2 h have h₂ : ¬ (d : ℤ) < 0 := of_decide_eq_false rfl - simp [h₁, h₂, ←Rat.normalize_eq_mk', Rat.normalize_eq_mkRat, ← mkRat_one, + simp [h₁, h₂, ← Rat.normalize_eq_mk', Rat.normalize_eq_mkRat, ← mkRat_one, Rat.mkRat_mul_mkRat] end Rat diff --git a/Mathlib/Data/Rat/Lemmas.lean b/Mathlib/Data/Rat/Lemmas.lean index d5ec923a04d42..7a9068ab26419 100644 --- a/Mathlib/Data/Rat/Lemmas.lean +++ b/Mathlib/Data/Rat/Lemmas.lean @@ -61,7 +61,7 @@ theorem num_den_mk {q : ℚ} {n d : ℤ} (hd : d ≠ 0) (qdf : q = n /. d) : theorem num_mk (n d : ℤ) : (n /. d).num = d.sign * n / n.gcd d := by rcases d with ((_ | _) | _) <;> - rw [←Int.div_eq_ediv_of_dvd] <;> + rw [← Int.div_eq_ediv_of_dvd] <;> simp [divInt, mkRat, Rat.normalize, Nat.succPNat, Int.sign, Int.gcd, -Nat.cast_succ, Int.zero_ediv, Int.ofNat_dvd_left, Nat.gcd_dvd_left] #align rat.num_mk Rat.num_mk @@ -210,15 +210,15 @@ theorem den_div_cast_eq_one_iff (m n : ℤ) (hn : n ≠ 0) : ((m : ℚ) / n).den theorem num_div_eq_of_coprime {a b : ℤ} (hb0 : 0 < b) (h : Nat.Coprime a.natAbs b.natAbs) : (a / b : ℚ).num = a := by -- Porting note: was `lift b to ℕ using le_of_lt hb0` - rw [←Int.natAbs_of_nonneg hb0.le, ← Rat.divInt_eq_div, - ←mk_eq_divInt _ _ (Int.natAbs_ne_zero.mpr hb0.ne') h] + rw [← Int.natAbs_of_nonneg hb0.le, ← Rat.divInt_eq_div, + ← mk_eq_divInt _ _ (Int.natAbs_ne_zero.mpr hb0.ne') h] #align rat.num_div_eq_of_coprime Rat.num_div_eq_of_coprime theorem den_div_eq_of_coprime {a b : ℤ} (hb0 : 0 < b) (h : Nat.Coprime a.natAbs b.natAbs) : ((a / b : ℚ).den : ℤ) = b := by -- Porting note: was `lift b to ℕ using le_of_lt hb0` - rw [←Int.natAbs_of_nonneg hb0.le, ← Rat.divInt_eq_div, - ←mk_eq_divInt _ _ (Int.natAbs_ne_zero.mpr hb0.ne') h] + rw [← Int.natAbs_of_nonneg hb0.le, ← Rat.divInt_eq_div, + ← mk_eq_divInt _ _ (Int.natAbs_ne_zero.mpr hb0.ne') h] #align rat.denom_div_eq_of_coprime Rat.den_div_eq_of_coprime theorem div_int_inj {a b c d : ℤ} (hb0 : 0 < b) (hd0 : 0 < d) (h1 : Nat.Coprime a.natAbs b.natAbs) diff --git a/Mathlib/Data/Rat/Order.lean b/Mathlib/Data/Rat/Order.lean index 4fd8fa956cdf2..c818c84d65440 100644 --- a/Mathlib/Data/Rat/Order.lean +++ b/Mathlib/Data/Rat/Order.lean @@ -102,7 +102,7 @@ def numDenCasesOn''.{u} {C : ℚ → Sort u} (a : ℚ) (H : ∀ (n : ℤ) (d : ℕ) (nz red), C (mk' n d nz red)) : C a := numDenCasesOn a fun n d h h' => by - rw [←mk_eq_divInt _ _ h.ne' h'] + rw [← mk_eq_divInt _ _ h.ne' h'] exact H n d h.ne' _ -- Porting note: TODO can this be shortened? @@ -126,7 +126,7 @@ protected theorem le_iff_Nonneg (a b : ℚ) : a ≤ b ↔ Rat.Nonneg (b - a) := · simp only [Nat.cast_pos] apply Nat.gcd_pos_of_pos_right apply mul_pos <;> rwa [pos_iff_ne_zero] - · simp only [divInt_ofNat, ←zero_iff_num_zero, mkRat_eq_zero hb] at h' + · simp only [divInt_ofNat, ← zero_iff_num_zero, mkRat_eq_zero hb] at h' simp [h', Rat.Nonneg] · simp [Rat.Nonneg, Rat.sub_def, normalize_eq] refine ⟨fun H => ?_, fun H _ => ?_⟩ @@ -141,7 +141,7 @@ protected theorem le_iff_Nonneg (a b : ℚ) : a ≤ b ↔ Rat.Nonneg (b - a) := · apply le_trans <| mul_nonpos_of_nonpos_of_nonneg ha (Nat.cast_nonneg _) exact mul_nonneg hb.le (Nat.cast_nonneg _) · exact H (fun _ => ha) - · rw [←sub_nonneg] + · rw [← sub_nonneg] contrapose! H apply Int.ediv_neg' H simp only [Nat.cast_pos] @@ -172,7 +172,7 @@ protected theorem le_total : a ≤ b ∨ b ≤ a := by protected theorem le_antisymm {a b : ℚ} (hab : a ≤ b) (hba : b ≤ a) : a = b := by rw [Rat.le_iff_Nonneg] at hab hba rw [sub_eq_add_neg] at hba - rw [←neg_sub, sub_eq_add_neg] at hab + rw [← neg_sub, sub_eq_add_neg] at hab have := eq_neg_of_add_eq_zero_left (Rat.nonneg_antisymm hba hab) rwa [neg_neg] at this #align rat.le_antisymm Rat.le_antisymm diff --git a/Mathlib/Data/Rat/Sqrt.lean b/Mathlib/Data/Rat/Sqrt.lean index 9e7dea9f265a3..8bcb2b93012e3 100644 --- a/Mathlib/Data/Rat/Sqrt.lean +++ b/Mathlib/Data/Rat/Sqrt.lean @@ -45,6 +45,6 @@ theorem sqrt_nonneg (q : ℚ) : 0 ≤ Rat.sqrt q := /-- `IsSquare` can be decided on `ℚ` by checking against the square root. -/ instance : DecidablePred (IsSquare : ℚ → Prop) := fun m => decidable_of_iff' (sqrt m * sqrt m = m) <| by - simp_rw [←exists_mul_self m, IsSquare, eq_comm] + simp_rw [← exists_mul_self m, IsSquare, eq_comm] end Rat diff --git a/Mathlib/Data/Real/ENNReal.lean b/Mathlib/Data/Real/ENNReal.lean index 8f1efa6147bab..352ea3fd73ba9 100644 --- a/Mathlib/Data/Real/ENNReal.lean +++ b/Mathlib/Data/Real/ENNReal.lean @@ -954,7 +954,7 @@ theorem coe_max (r p : ℝ≥0) : ((max r p : ℝ≥0) : ℝ≥0∞) = max (r : theorem le_of_top_imp_top_of_toNNReal_le {a b : ℝ≥0∞} (h : a = ⊤ → b = ⊤) (h_nnreal : a ≠ ⊤ → b ≠ ⊤ → a.toNNReal ≤ b.toNNReal) : a ≤ b := by - by_contra' hlt + by_contra! hlt lift b to ℝ≥0 using hlt.ne_top lift a to ℝ≥0 using mt h coe_ne_top refine hlt.not_le ?_ @@ -1442,7 +1442,7 @@ protected theorem mul_div_cancel' (h0 : a ≠ 0) (hI : a ≠ ∞) : a * (b / a) -- porting note: `simp only [div_eq_mul_inv, mul_comm, mul_assoc]` doesn't work in the following two protected theorem mul_comm_div : a / b * c = a * (c / b) := by - simp only [div_eq_mul_inv, mul_right_comm, ←mul_assoc] + simp only [div_eq_mul_inv, mul_right_comm, ← mul_assoc] #align ennreal.mul_comm_div ENNReal.mul_comm_div protected theorem mul_div_right_comm : a * b / c = a / c * b := by @@ -2362,7 +2362,7 @@ theorem toNNReal_pow (a : ℝ≥0∞) (n : ℕ) : (a ^ n).toNNReal = a.toNNReal @[simp] theorem toNNReal_prod {ι : Type*} {s : Finset ι} {f : ι → ℝ≥0∞} : (∏ i in s, f i).toNNReal = ∏ i in s, (f i).toNNReal := - toNNRealHom.map_prod _ _ + map_prod toNNRealHom _ _ #align ennreal.to_nnreal_prod ENNReal.toNNReal_prod -- porting note: todo: upgrade to `→*₀` @@ -2386,7 +2386,7 @@ theorem toReal_pow (a : ℝ≥0∞) (n : ℕ) : (a ^ n).toReal = a.toReal ^ n := @[simp] theorem toReal_prod {ι : Type*} {s : Finset ι} {f : ι → ℝ≥0∞} : (∏ i in s, f i).toReal = ∏ i in s, (f i).toReal := - toRealHom.map_prod _ _ + map_prod toRealHom _ _ #align ennreal.to_real_prod ENNReal.toReal_prod theorem toReal_ofReal_mul (c : ℝ) (a : ℝ≥0∞) (h : 0 ≤ c) : diff --git a/Mathlib/Data/Real/ENatENNReal.lean b/Mathlib/Data/Real/ENatENNReal.lean index 0780cca8cc8c9..65f8754e5ee75 100644 --- a/Mathlib/Data/Real/ENatENNReal.lean +++ b/Mathlib/Data/Real/ENatENNReal.lean @@ -56,8 +56,10 @@ theorem toENNReal_coe (n : ℕ) : ((n : ℕ∞) : ℝ≥0∞) = n := rfl #align enat.coe_ennreal_coe ENat.toENNReal_coe +-- See note [no_index around OfNat.ofNat] @[simp, norm_cast] -theorem toENNReal_ofNat (n : ℕ) [n.AtLeastTwo] : ((OfNat.ofNat n : ℕ∞) : ℝ≥0∞) = OfNat.ofNat n := +theorem toENNReal_ofNat (n : ℕ) [n.AtLeastTwo] : + ((no_index (OfNat.ofNat n : ℕ∞)) : ℝ≥0∞) = OfNat.ofNat n := rfl @[simp, norm_cast] diff --git a/Mathlib/Data/Real/Hyperreal.lean b/Mathlib/Data/Real/Hyperreal.lean index 06dad5c7ef25c..b7508f9a63eec 100644 --- a/Mathlib/Data/Real/Hyperreal.lean +++ b/Mathlib/Data/Real/Hyperreal.lean @@ -89,8 +89,11 @@ theorem coe_add (x y : ℝ) : ↑(x + y) = (x + y : ℝ*) := #noalign hyperreal.coe_bit0 #noalign hyperreal.coe_bit1 +-- See note [no_index around OfNat.ofNat] @[simp, norm_cast] -theorem coe_ofNat (n : ℕ) [n.AtLeastTwo] : ((OfNat.ofNat n : ℝ) : ℝ*) = OfNat.ofNat n := rfl +theorem coe_ofNat (n : ℕ) [n.AtLeastTwo] : + ((no_index (OfNat.ofNat n : ℝ)) : ℝ*) = OfNat.ofNat n := + rfl @[simp, norm_cast] theorem coe_mul (x y : ℝ) : ↑(x * y) = (x * y : ℝ*) := diff --git a/Mathlib/Data/Rel.lean b/Mathlib/Data/Rel.lean index e004ce8a9d1d1..f3e79b58f0843 100644 --- a/Mathlib/Data/Rel.lean +++ b/Mathlib/Data/Rel.lean @@ -291,7 +291,7 @@ theorem image_inter_dom_eq (s : Set α) : r.image (s ∩ r.dom) = r.image s := b @[simp] theorem preimage_inter_codom_eq (s : Set β) : r.preimage (s ∩ r.codom) = r.preimage s := by - rw[←dom_inv, preimage, preimage, image_inter_dom_eq] + rw[← dom_inv, preimage, preimage, image_inter_dom_eq] theorem inter_dom_subset_preimage_image (s : Set α) : s ∩ r.dom ⊆ r.preimage (r.image s) := by intro x hx diff --git a/Mathlib/Data/Semiquot.lean b/Mathlib/Data/Semiquot.lean index cfae2cb5544ad..0060288c10d25 100644 --- a/Mathlib/Data/Semiquot.lean +++ b/Mathlib/Data/Semiquot.lean @@ -226,7 +226,7 @@ theorem eq_pure {q : Semiquot α} (p) : q = pure (get q p) := theorem pure_isPure (a : α) : IsPure (pure a) | b, ab, c, ac => by rw [mem_pure] at ab ac - rwa [←ac] at ab + rwa [← ac] at ab #align semiquot.pure_is_pure Semiquot.pure_isPure theorem isPure_iff {s : Semiquot α} : IsPure s ↔ ∃ a, s = pure a := diff --git a/Mathlib/Data/Seq/WSeq.lean b/Mathlib/Data/Seq/WSeq.lean index 657fe4156daf1..7419c171fb53b 100644 --- a/Mathlib/Data/Seq/WSeq.lean +++ b/Mathlib/Data/Seq/WSeq.lean @@ -721,7 +721,7 @@ theorem dropn_cons (a : α) (s) (n) : drop (cons a s) (n + 1) = drop s n := by | zero => simp [drop] | succ n n_ih => -- Porting note: Was `simp [*, drop]`. - simp [drop, ←n_ih] + simp [drop, ← n_ih] #align stream.wseq.dropn_cons Stream'.WSeq.dropn_cons @[simp] diff --git a/Mathlib/Data/Set/Basic.lean b/Mathlib/Data/Set/Basic.lean index 1422056623191..333aed946628a 100644 --- a/Mathlib/Data/Set/Basic.lean +++ b/Mathlib/Data/Set/Basic.lean @@ -2524,7 +2524,7 @@ theorem nontrivial_of_exists_ne {x} (hx : x ∈ s) (h : ∃ y ∈ s, y ≠ x) : #align set.nontrivial_of_exists_ne Set.nontrivial_of_exists_ne theorem Nontrivial.exists_ne (hs : s.Nontrivial) (z) : ∃ x ∈ s, x ≠ z := by - by_contra' H + by_contra! H rcases hs with ⟨x, hx, y, hy, hxy⟩ rw [H x hx, H y hy] at hxy exact hxy rfl @@ -2659,7 +2659,7 @@ protected lemma subsingleton_or_nontrivial (s : Set α) : s.Subsingleton ∨ s.N #align set.subsingleton_or_nontrivial Set.subsingleton_or_nontrivial lemma eq_singleton_or_nontrivial (ha : a ∈ s) : s = {a} ∨ s.Nontrivial := by - rw [←subsingleton_iff_singleton ha]; exact s.subsingleton_or_nontrivial + rw [← subsingleton_iff_singleton ha]; exact s.subsingleton_or_nontrivial #align set.eq_singleton_or_nontrivial Set.eq_singleton_or_nontrivial lemma nontrivial_iff_ne_singleton (ha : a ∈ s) : s.Nontrivial ↔ s ≠ {a} := diff --git a/Mathlib/Data/Set/Card.lean b/Mathlib/Data/Set/Card.lean index 12f604840ca64..d61cf4bc106c3 100644 --- a/Mathlib/Data/Set/Card.lean +++ b/Mathlib/Data/Set/Card.lean @@ -86,11 +86,11 @@ theorem encard_coe_eq_coe_finsetCard (s : Finset α) : encard (s : Set α) = s.c theorem Infinite.encard_eq {s : Set α} (h : s.Infinite) : s.encard = ⊤ := by have := h.to_subtype - rw [encard, ←PartENat.withTopEquiv.symm.injective.eq_iff, Equiv.symm_apply_apply, + rw [encard, ← PartENat.withTopEquiv.symm.injective.eq_iff, Equiv.symm_apply_apply, PartENat.withTopEquiv_symm_top, PartENat.card_eq_top_of_infinite] @[simp] theorem encard_eq_zero : s.encard = 0 ↔ s = ∅ := by - rw [encard, ←PartENat.withTopEquiv.symm.injective.eq_iff, Equiv.symm_apply_apply, + rw [encard, ← PartENat.withTopEquiv.symm.injective.eq_iff, Equiv.symm_apply_apply, PartENat.withTopEquiv_symm_zero, PartENat.card_eq_zero_iff_empty, isEmpty_subtype, eq_empty_iff_forall_not_mem] @@ -98,7 +98,7 @@ theorem Infinite.encard_eq {s : Set α} (h : s.Infinite) : s.encard = ⊤ := by rw [encard_eq_zero] theorem nonempty_of_encard_ne_zero (h : s.encard ≠ 0) : s.Nonempty := by - rwa [nonempty_iff_ne_empty, Ne.def, ←encard_eq_zero] + rwa [nonempty_iff_ne_empty, Ne.def, ← encard_eq_zero] theorem encard_ne_zero : s.encard ≠ 0 ↔ s.Nonempty := by rw [ne_eq, encard_eq_zero, nonempty_iff_ne_empty] @@ -107,16 +107,16 @@ theorem encard_ne_zero : s.encard ≠ 0 ↔ s.Nonempty := by rw [pos_iff_ne_zero, encard_ne_zero] @[simp] theorem encard_singleton (e : α) : ({e} : Set α).encard = 1 := by - rw [encard, ←PartENat.withTopEquiv.symm.injective.eq_iff, Equiv.symm_apply_apply, + rw [encard, ← PartENat.withTopEquiv.symm.injective.eq_iff, Equiv.symm_apply_apply, PartENat.card_eq_coe_fintype_card, Fintype.card_ofSubsingleton, Nat.cast_one]; rfl theorem encard_union_eq (h : Disjoint s t) : (s ∪ t).encard = s.encard + t.encard := by classical - have e := (Equiv.Set.union (by rwa [subset_empty_iff, ←disjoint_iff_inter_eq_empty])).symm - simp [encard, ←PartENat.card_congr e, PartENat.card_sum, PartENat.withTopEquiv] + have e := (Equiv.Set.union (by rwa [subset_empty_iff, ← disjoint_iff_inter_eq_empty])).symm + simp [encard, ← PartENat.card_congr e, PartENat.card_sum, PartENat.withTopEquiv] theorem encard_insert_of_not_mem (has : a ∉ s) : (insert a s).encard = s.encard + 1 := by - rw [←union_singleton, encard_union_eq (by simpa), encard_singleton] + rw [← union_singleton, encard_union_eq (by simpa), encard_singleton] theorem Finite.encard_lt_top (h : s.Finite) : s.encard < ⊤ := by refine' h.induction_on (by simp) _ @@ -134,13 +134,13 @@ theorem Finite.exists_encard_eq_coe (h : s.Finite) : ∃ (n : ℕ), s.encard = n ⟨fun h ↦ by_contra fun h' ↦ h.ne (Infinite.encard_eq h'), Finite.encard_lt_top⟩ @[simp] theorem encard_eq_top_iff : s.encard = ⊤ ↔ s.Infinite := by - rw [←not_iff_not, ←Ne.def, ←lt_top_iff_ne_top, encard_lt_top_iff, not_infinite] + rw [← not_iff_not, ← Ne.def, ← lt_top_iff_ne_top, encard_lt_top_iff, not_infinite] theorem encard_ne_top_iff : s.encard ≠ ⊤ ↔ s.Finite := by simp theorem finite_of_encard_le_coe {k : ℕ} (h : s.encard ≤ k) : s.Finite := by - rw [←encard_lt_top_iff]; exact h.trans_lt (WithTop.coe_lt_top _) + rw [← encard_lt_top_iff]; exact h.trans_lt (WithTop.coe_lt_top _) theorem finite_of_encard_eq_coe {k : ℕ} (h : s.encard = k) : s.Finite := finite_of_encard_le_coe h.le @@ -152,25 +152,25 @@ theorem encard_le_coe_iff {k : ℕ} : s.encard ≤ k ↔ s.Finite ∧ ∃ (n₀ section Lattice theorem encard_le_of_subset (h : s ⊆ t) : s.encard ≤ t.encard := by - rw [←union_diff_cancel h, encard_union_eq disjoint_sdiff_right]; exact le_self_add + rw [← union_diff_cancel h, encard_union_eq disjoint_sdiff_right]; exact le_self_add theorem encard_mono {α : Type*} : Monotone (encard : Set α → ℕ∞) := fun _ _ ↦ encard_le_of_subset theorem encard_diff_add_encard_of_subset (h : s ⊆ t) : (t \ s).encard + s.encard = t.encard := by - rw [←encard_union_eq disjoint_sdiff_left, diff_union_self, union_eq_self_of_subset_right h] + rw [← encard_union_eq disjoint_sdiff_left, diff_union_self, union_eq_self_of_subset_right h] @[simp] theorem one_le_encard_iff_nonempty : 1 ≤ s.encard ↔ s.Nonempty := by - rw [nonempty_iff_ne_empty, Ne.def, ←encard_eq_zero, ENat.one_le_iff_ne_zero] + rw [nonempty_iff_ne_empty, Ne.def, ← encard_eq_zero, ENat.one_le_iff_ne_zero] theorem encard_diff_add_encard_inter (s t : Set α) : (s \ t).encard + (s ∩ t).encard = s.encard := by - rw [←encard_union_eq (disjoint_of_subset_right (inter_subset_right _ _) disjoint_sdiff_left), + rw [← encard_union_eq (disjoint_of_subset_right (inter_subset_right _ _) disjoint_sdiff_left), diff_union_inter] theorem encard_union_add_encard_inter (s t : Set α) : (s ∪ t).encard + (s ∩ t).encard = s.encard + t.encard := -by rw [←diff_union_self, encard_union_eq disjoint_sdiff_left, add_right_comm, +by rw [← diff_union_self, encard_union_eq disjoint_sdiff_left, add_right_comm, encard_diff_add_encard_inter] theorem encard_eq_encard_iff_encard_diff_eq_encard_diff (h : (s ∩ t).Finite) : @@ -189,13 +189,13 @@ theorem encard_lt_encard_iff_encard_diff_lt_encard_diff (h : (s ∩ t).Finite) : WithTop.add_lt_add_iff_right h.encard_lt_top.ne] theorem encard_union_le (s t : Set α) : (s ∪ t).encard ≤ s.encard + t.encard := by - rw [←encard_union_add_encard_inter]; exact le_self_add + rw [← encard_union_add_encard_inter]; exact le_self_add theorem finite_iff_finite_of_encard_eq_encard (h : s.encard = t.encard) : s.Finite ↔ t.Finite := by - rw [←encard_lt_top_iff, ←encard_lt_top_iff, h] + rw [← encard_lt_top_iff, ← encard_lt_top_iff, h] theorem infinite_iff_infinite_of_encard_eq_encard (h : s.encard = t.encard) : - s.Infinite ↔ t.Infinite := by rw [←encard_eq_top_iff, h, encard_eq_top_iff] + s.Infinite ↔ t.Infinite := by rw [← encard_eq_top_iff, h, encard_eq_top_iff] theorem Finite.finite_of_encard_le {s : Set α} {t : Set β} (hs : s.Finite) (h : t.encard ≤ s.encard) : t.Finite := @@ -203,7 +203,7 @@ theorem Finite.finite_of_encard_le {s : Set α} {t : Set β} (hs : s.Finite) theorem Finite.eq_of_subset_of_encard_le (ht : t.Finite) (hst : s ⊆ t) (hts : t.encard ≤ s.encard) : s = t := by - rw [←zero_add (a := encard s), ←encard_diff_add_encard_of_subset hst] at hts + rw [← zero_add (a := encard s), ← encard_diff_add_encard_of_subset hst] at hts have hdiff := WithTop.le_of_add_le_add_right (ht.subset hst).encard_lt_top.ne hts rw [nonpos_iff_eq_zero, encard_eq_zero, diff_eq_empty] at hdiff exact hst.antisymm hdiff @@ -219,7 +219,7 @@ theorem encard_strictMono [Finite α] : StrictMono (encard : Set α → ℕ∞) fun _ _ h ↦ (toFinite _).encard_lt_encard h theorem encard_diff_add_encard (s t : Set α) : (s \ t).encard + t.encard = (s ∪ t).encard := by - rw [←encard_union_eq disjoint_sdiff_left, diff_union_self] + rw [← encard_union_eq disjoint_sdiff_left, diff_union_self] theorem encard_le_encard_diff_add_encard (s t : Set α) : s.encard ≤ (s \ t).encard + t.encard := (encard_mono (subset_union_left s t)).trans_eq (encard_diff_add_encard _ _).symm @@ -228,44 +228,44 @@ theorem tsub_encard_le_encard_diff (s t : Set α) : s.encard - t.encard ≤ (s \ rw [tsub_le_iff_left, add_comm]; apply encard_le_encard_diff_add_encard theorem encard_add_encard_compl (s : Set α) : s.encard + sᶜ.encard = (univ : Set α).encard := by - rw [←encard_union_eq disjoint_compl_right, union_compl_self] + rw [← encard_union_eq disjoint_compl_right, union_compl_self] end Lattice section InsertErase theorem encard_insert_le (s : Set α) (x : α) : (insert x s).encard ≤ s.encard + 1 := by - rw [←union_singleton, ←encard_singleton x]; apply encard_union_le + rw [← union_singleton, ← encard_singleton x]; apply encard_union_le theorem encard_singleton_inter (s : Set α) (x : α) : ({x} ∩ s).encard ≤ 1 := by - rw [←encard_singleton x]; exact encard_le_of_subset (inter_subset_left _ _) + rw [← encard_singleton x]; exact encard_le_of_subset (inter_subset_left _ _) theorem encard_diff_singleton_add_one (h : a ∈ s) : (s \ {a}).encard + 1 = s.encard := by - rw [←encard_insert_of_not_mem (fun h ↦ h.2 rfl), insert_diff_singleton, insert_eq_of_mem h] + rw [← encard_insert_of_not_mem (fun h ↦ h.2 rfl), insert_diff_singleton, insert_eq_of_mem h] theorem encard_diff_singleton_of_mem (h : a ∈ s) : (s \ {a}).encard = s.encard - 1 := by - rw [←encard_diff_singleton_add_one h, ←WithTop.add_right_cancel_iff WithTop.one_ne_top, + rw [← encard_diff_singleton_add_one h, ← WithTop.add_right_cancel_iff WithTop.one_ne_top, tsub_add_cancel_of_le (self_le_add_left _ _)] theorem encard_tsub_one_le_encard_diff_singleton (s : Set α) (x : α) : s.encard - 1 ≤ (s \ {x}).encard := by - rw [←encard_singleton x]; apply tsub_encard_le_encard_diff + rw [← encard_singleton x]; apply tsub_encard_le_encard_diff theorem encard_exchange (ha : a ∉ s) (hb : b ∈ s) : (insert a (s \ {b})).encard = s.encard := by rw [encard_insert_of_not_mem, encard_diff_singleton_add_one hb] simp_all only [not_true, mem_diff, mem_singleton_iff, false_and, not_false_eq_true] theorem encard_exchange' (ha : a ∉ s) (hb : b ∈ s) : (insert a s \ {b}).encard = s.encard := by - rw [←insert_diff_singleton_comm (by rintro rfl; exact ha hb), encard_exchange ha hb] + rw [← insert_diff_singleton_comm (by rintro rfl; exact ha hb), encard_exchange ha hb] theorem encard_eq_add_one_iff {k : ℕ∞} : s.encard = k + 1 ↔ (∃ a t, ¬a ∈ t ∧ insert a t = s ∧ t.encard = k) := by refine' ⟨fun h ↦ _, _⟩ · obtain ⟨a, ha⟩ := nonempty_of_encard_ne_zero (s := s) (by simp [h]) refine' ⟨a, s \ {a}, fun h ↦ h.2 rfl, by rwa [insert_diff_singleton, insert_eq_of_mem], _⟩ - rw [←WithTop.add_right_cancel_iff WithTop.one_ne_top, ←h, + rw [← WithTop.add_right_cancel_iff WithTop.one_ne_top, ← h, encard_diff_singleton_add_one ha] rintro ⟨a, t, h, rfl, rfl⟩ rw [encard_insert_of_not_mem h] @@ -276,7 +276,7 @@ theorem eq_empty_or_encard_eq_top_or_encard_diff_singleton_lt (s : Set α) : s = ∅ ∨ s.encard = ⊤ ∨ ∃ a ∈ s, (s \ {a}).encard < s.encard := by refine' s.eq_empty_or_nonempty.elim Or.inl (Or.inr ∘ fun ⟨a,ha⟩ ↦ (s.finite_or_infinite.elim (fun hfin ↦ Or.inr ⟨a, ha, _⟩) (Or.inl ∘ Infinite.encard_eq))) - rw [←encard_diff_singleton_add_one ha]; nth_rw 1 [←add_zero (encard _)] + rw [← encard_diff_singleton_add_one ha]; nth_rw 1 [← add_zero (encard _)] exact WithTop.add_lt_add_left (hfin.diff _).encard_lt_top.ne zero_lt_one end InsertErase @@ -284,7 +284,7 @@ end InsertErase section SmallSets theorem encard_pair (hne : x ≠ y) : ({x, y} : Set α).encard = 2 := by - rw [encard_insert_of_not_mem (by simpa), ←one_add_one_eq_two, + rw [encard_insert_of_not_mem (by simpa), ← one_add_one_eq_two, WithTop.add_right_cancel_iff WithTop.one_ne_top, encard_singleton] theorem encard_eq_one : s.encard = 1 ↔ ∃ x, s = {x} := by @@ -297,17 +297,17 @@ theorem encard_le_one_iff_eq : s.encard ≤ 1 ↔ s = ∅ ∨ ∃ x, s = {x} := encard_eq_one] theorem encard_le_one_iff : s.encard ≤ 1 ↔ ∀ a b, a ∈ s → b ∈ s → a = b := by - rw [encard_le_one_iff_eq, or_iff_not_imp_left, ←Ne.def, ←nonempty_iff_ne_empty] + rw [encard_le_one_iff_eq, or_iff_not_imp_left, ← Ne.def, ← nonempty_iff_ne_empty] refine' ⟨fun h a b has hbs ↦ _, fun h ⟨x, hx⟩ ↦ ⟨x, ((singleton_subset_iff.2 hx).antisymm' (fun y hy ↦ h _ _ hy hx))⟩⟩ obtain ⟨x, rfl⟩ := h ⟨_, has⟩ rw [(has : a = x), (hbs : b = x)] theorem one_lt_encard_iff : 1 < s.encard ↔ ∃ a b, a ∈ s ∧ b ∈ s ∧ a ≠ b := by - rw [←not_iff_not, not_exists, not_lt, encard_le_one_iff]; aesop + rw [← not_iff_not, not_exists, not_lt, encard_le_one_iff]; aesop theorem exists_ne_of_one_lt_encard (h : 1 < s.encard) (a : α) : ∃ b ∈ s, b ≠ a := by - by_contra' h' + by_contra! h' obtain ⟨b, b', hb, hb', hne⟩ := one_lt_encard_iff.1 h apply hne rw [h' b hb, h' b' hb'] @@ -315,24 +315,24 @@ theorem exists_ne_of_one_lt_encard (h : 1 < s.encard) (a : α) : ∃ b ∈ s, b theorem encard_eq_two : s.encard = 2 ↔ ∃ x y, x ≠ y ∧ s = {x, y} := by refine' ⟨fun h ↦ _, fun ⟨x, y, hne, hs⟩ ↦ by rw [hs, encard_pair hne]⟩ obtain ⟨x, hx⟩ := nonempty_of_encard_ne_zero (s := s) (by rw [h]; simp) - rw [←insert_eq_of_mem hx, ←insert_diff_singleton, encard_insert_of_not_mem (fun h ↦ h.2 rfl), - ←one_add_one_eq_two, WithTop.add_right_cancel_iff (WithTop.one_ne_top), encard_eq_one] at h + rw [← insert_eq_of_mem hx, ← insert_diff_singleton, encard_insert_of_not_mem (fun h ↦ h.2 rfl), + ← one_add_one_eq_two, WithTop.add_right_cancel_iff (WithTop.one_ne_top), encard_eq_one] at h obtain ⟨y, h⟩ := h refine' ⟨x, y, by rintro rfl; exact (h.symm.subset rfl).2 rfl, _⟩ - rw [←h, insert_diff_singleton, insert_eq_of_mem hx] + rw [← h, insert_diff_singleton, insert_eq_of_mem hx] theorem encard_eq_three {α : Type u_1} {s : Set α} : encard s = 3 ↔ ∃ x y z, x ≠ y ∧ x ≠ z ∧ y ≠ z ∧ s = {x, y, z} := by refine' ⟨fun h ↦ _, fun ⟨x, y, z, hxy, hyz, hxz, hs⟩ ↦ _⟩ · obtain ⟨x, hx⟩ := nonempty_of_encard_ne_zero (s := s) (by rw [h]; simp) - rw [←insert_eq_of_mem hx, ←insert_diff_singleton, + rw [← insert_eq_of_mem hx, ← insert_diff_singleton, encard_insert_of_not_mem (fun h ↦ h.2 rfl), (by exact rfl : (3 : ℕ∞) = 2 + 1), WithTop.add_right_cancel_iff WithTop.one_ne_top, encard_eq_two] at h obtain ⟨y, z, hne, hs⟩ := h refine' ⟨x, y, z, _, _, hne, _⟩ · rintro rfl; exact (hs.symm.subset (Or.inl rfl)).2 rfl · rintro rfl; exact (hs.symm.subset (Or.inr rfl)).2 rfl - rw [←hs, insert_diff_singleton, insert_eq_of_mem hx] + rw [← hs, insert_diff_singleton, insert_eq_of_mem hx] rw [hs, encard_insert_of_not_mem, encard_insert_of_not_mem, encard_singleton] <;> aesop theorem Nat.encard_range (k : ℕ) : {i | i < k}.encard = k := by @@ -344,9 +344,9 @@ end SmallSets theorem Finite.eq_insert_of_subset_of_encard_eq_succ (hs : s.Finite) (h : s ⊆ t) (hst : t.encard = s.encard + 1) : ∃ a, t = insert a s := by - rw [←encard_diff_add_encard_of_subset h, add_comm, + rw [← encard_diff_add_encard_of_subset h, add_comm, WithTop.add_left_cancel_iff hs.encard_lt_top.ne, encard_eq_one] at hst - obtain ⟨x, hx⟩ := hst; use x; rw [←diff_union_of_subset h, hx, singleton_union] + obtain ⟨x, hx⟩ := hst; use x; rw [← diff_union_of_subset h, hx, singleton_union] theorem exists_subset_encard_eq (hk : k ≤ s.encard) : ∃ t, t ⊆ s ∧ t.encard = k := by revert hk @@ -354,7 +354,7 @@ theorem exists_subset_encard_eq (hk : k ≤ s.encard) : ∃ t, t ⊆ s ∧ t.enc · obtain ⟨t₀, ht₀s, ht₀⟩ := IH (le_trans (by simp) hle) simp only [Nat.cast_succ] at * have hne : t₀ ≠ s - · rintro rfl; rw [ht₀, ←Nat.cast_one, ←Nat.cast_add, Nat.cast_le] at hle; simp at hle + · rintro rfl; rw [ht₀, ← Nat.cast_one, ← Nat.cast_add, Nat.cast_le] at hle; simp at hle obtain ⟨x, hx⟩ := exists_of_ssubset (ht₀s.ssubset_of_ne hne) exact ⟨insert x t₀, insert_subset hx.1 ht₀s, by rw [encard_insert_of_not_mem hx.2, ht₀]⟩ simp only [top_le_iff, encard_eq_top_iff] @@ -367,7 +367,7 @@ theorem exists_superset_subset_encard_eq (hst : s ⊆ t) (hsk : s.encard ≤ k) obtain ⟨k, rfl⟩ := exists_add_of_le hsk obtain ⟨k', hk'⟩ := exists_add_of_le hkt have hk : k ≤ encard (t \ s) - · rw [←encard_diff_add_encard_of_subset hst, add_comm] at hkt + · rw [← encard_diff_add_encard_of_subset hst, add_comm] at hkt exact WithTop.le_of_add_le_add_right hs hkt obtain ⟨r', hr', rfl⟩ := exists_subset_encard_eq hk refine' ⟨s ∪ r', subset_union_left _ _, union_subset hst (hr'.trans (diff_subset _ _)), _⟩ @@ -381,20 +381,20 @@ theorem InjOn.encard_image (h : InjOn f s) : (f '' s).encard = s.encard := by rw [encard, PartENat.card_image_of_injOn h, encard] theorem encard_congr (e : s ≃ t) : s.encard = t.encard := by - rw [←encard_univ_coe, ←encard_univ_coe t, encard_univ, encard_univ, PartENat.card_congr e] + rw [← encard_univ_coe, ← encard_univ_coe t, encard_univ, encard_univ, PartENat.card_congr e] theorem _root_.Function.Injective.encard_image (hf : f.Injective) (s : Set α) : (f '' s).encard = s.encard := (hf.injOn s).encard_image theorem _root_.Function.Embedding.enccard_le (e : s ↪ t) : s.encard ≤ t.encard := by - rw [←encard_univ_coe, ←e.injective.encard_image, ←Subtype.coe_injective.encard_image] + rw [← encard_univ_coe, ← e.injective.encard_image, ← Subtype.coe_injective.encard_image] exact encard_mono (by simp) theorem encard_image_le (f : α → β) (s : Set α) : (f '' s).encard ≤ s.encard := by obtain (h | h) := isEmpty_or_nonempty α · rw [s.eq_empty_of_isEmpty]; simp - rw [←(f.invFunOn_injOn_image s).encard_image] + rw [← (f.invFunOn_injOn_image s).encard_image] apply encard_le_of_subset exact f.invFunOn_image_image_subset s @@ -402,17 +402,17 @@ theorem Finite.injOn_of_encard_image_eq (hs : s.Finite) (h : (f '' s).encard = s InjOn f s := by obtain (h' | hne) := isEmpty_or_nonempty α · rw [s.eq_empty_of_isEmpty]; simp - rw [←(f.invFunOn_injOn_image s).encard_image] at h + rw [← (f.invFunOn_injOn_image s).encard_image] at h rw [injOn_iff_invFunOn_image_image_eq_self] exact hs.eq_of_subset_of_encard_le (f.invFunOn_image_image_subset s) h.symm.le theorem encard_preimage_of_injective_subset_range (hf : f.Injective) (ht : t ⊆ range f) : (f ⁻¹' t).encard = t.encard := by - rw [←hf.encard_image, image_preimage_eq_inter_range, inter_eq_self_of_subset_left ht] + rw [← hf.encard_image, image_preimage_eq_inter_range, inter_eq_self_of_subset_left ht] theorem encard_le_encard_of_injOn (hf : MapsTo f s t) (f_inj : InjOn f s) : s.encard ≤ t.encard := by - rw [←f_inj.encard_image]; apply encard_le_of_subset; rintro _ ⟨x, hx, rfl⟩; exact hf hx + rw [← f_inj.encard_image]; apply encard_le_of_subset; rintro _ ⟨x, hx, rfl⟩; exact hf hx theorem Finite.exists_injOn_of_encard_le [Nonempty β] {s : Set α} {t : Set β} (hs : s.Finite) (hle : s.encard ≤ t.encard) : ∃ (f : α → β), s ⊆ f ⁻¹' t ∧ InjOn f s := by @@ -422,14 +422,14 @@ theorem Finite.exists_injOn_of_encard_le [Nonempty β] {s : Set α} {t : Set β} · exact (encard_ne_top_iff.mpr hs h).elim obtain ⟨b, hbt⟩ := encard_pos.1 ((encard_pos.2 ⟨_, has⟩).trans_le hle) have hle' : (s \ {a}).encard ≤ (t \ {b}).encard - · rwa [←WithTop.add_le_add_iff_right WithTop.one_ne_top, + · rwa [← WithTop.add_le_add_iff_right WithTop.one_ne_top, encard_diff_singleton_add_one has, encard_diff_singleton_add_one hbt] obtain ⟨f₀, hf₀s, hinj⟩ := exists_injOn_of_encard_le (hs.diff {a}) hle' simp only [preimage_diff, subset_def, mem_diff, mem_singleton_iff, mem_preimage, and_imp] at hf₀s use Function.update f₀ a b - rw [←insert_eq_of_mem has, ←insert_diff_singleton, injOn_insert (fun h ↦ h.2 rfl)] + rw [← insert_eq_of_mem has, ← insert_diff_singleton, injOn_insert (fun h ↦ h.2 rfl)] simp only [mem_diff, mem_singleton_iff, not_true, and_false, insert_diff_singleton, subset_def, mem_insert_iff, mem_preimage, ne_eq, Function.update_apply, forall_eq_or_imp, ite_true, and_imp, mem_image, ite_eq_left_iff, not_exists, not_and, not_forall, exists_prop, and_iff_right hbt] @@ -464,7 +464,7 @@ syntax "to_encard_tac" : tactic macro_rules | `(tactic| to_encard_tac) => `(tactic| - simp only [←Nat.cast_le (α := ℕ∞), ←Nat.cast_inj (R := ℕ∞), Nat.cast_add, Nat.cast_one]) + simp only [← Nat.cast_le (α := ℕ∞), ← Nat.cast_inj (R := ℕ∞), Nat.cast_add, Nat.cast_one]) /-- The cardinality of `s : Set α` . Has the junk value `0` if `s` is infinite -/ @@ -487,13 +487,13 @@ theorem Finite.cast_ncard_eq (hs : s.Finite) : s.ncard = s.encard := by theorem ncard_eq_toFinset_card (s : Set α) (hs : s.Finite := by toFinite_tac) : s.ncard = hs.toFinset.card := by - rw [←Nat.card_coe_set_eq, @Nat.card_eq_fintype_card _ hs.fintype, + rw [← Nat.card_coe_set_eq, @Nat.card_eq_fintype_card _ hs.fintype, @Finite.card_toFinset _ _ hs.fintype hs] #align set.ncard_eq_to_finset_card Set.ncard_eq_toFinset_card theorem ncard_eq_toFinset_card' (s : Set α) [Fintype s] : s.ncard = s.toFinset.card := by - simp [←Nat.card_coe_set_eq, Nat.card_eq_fintype_card] + simp [← Nat.card_coe_set_eq, Nat.card_eq_fintype_card] theorem encard_le_coe_iff_finite_ncard_le {k : ℕ} : s.encard ≤ k ↔ s.Finite ∧ s.ncard ≤ k := by rw [encard_le_coe_iff, and_congr_right_iff] @@ -501,12 +501,12 @@ theorem encard_le_coe_iff_finite_ncard_le {k : ℕ} : s.encard ≤ k ↔ s.Finit fun h ↦ ⟨s.ncard, by rw [hfin.cast_ncard_eq], h⟩⟩ theorem Infinite.ncard (hs : s.Infinite) : s.ncard = 0 := by - rw [←Nat.card_coe_set_eq, @Nat.card_eq_zero_of_infinite _ hs.to_subtype] + rw [← Nat.card_coe_set_eq, @Nat.card_eq_zero_of_infinite _ hs.to_subtype] #align set.infinite.ncard Set.Infinite.ncard theorem ncard_le_of_subset (hst : s ⊆ t) (ht : t.Finite := by toFinite_tac) : s.ncard ≤ t.ncard := by - rw [←Nat.cast_le (α := ℕ∞), ht.cast_ncard_eq, (ht.subset hst).cast_ncard_eq] + rw [← Nat.cast_le (α := ℕ∞), ht.cast_ncard_eq, (ht.subset hst).cast_ncard_eq] exact encard_mono hst #align set.ncard_le_of_subset Set.ncard_le_of_subset @@ -515,7 +515,7 @@ theorem ncard_mono [Finite α] : @Monotone (Set α) _ _ _ ncard := fun _ _ ↦ n @[simp] theorem ncard_eq_zero (hs : s.Finite := by toFinite_tac) : s.ncard = 0 ↔ s = ∅ := by - rw [←Nat.cast_inj (R := ℕ∞), hs.cast_ncard_eq, Nat.cast_zero, encard_eq_zero] + rw [← Nat.cast_inj (R := ℕ∞), hs.cast_ncard_eq, Nat.cast_zero, encard_eq_zero] #align set.ncard_eq_zero Set.ncard_eq_zero @[simp] theorem ncard_coe_Finset (s : Finset α) : (s : Set α).ncard = s.card := by @@ -559,14 +559,14 @@ theorem nonempty_of_ncard_ne_zero (hs : s.ncard ≠ 0) : s.Nonempty := by #align set.ncard_singleton Set.ncard_singleton theorem ncard_singleton_inter (a : α) (s : Set α) : ({a} ∩ s).ncard ≤ 1 := by - rw [←Nat.cast_le (α := ℕ∞), (toFinite _).cast_ncard_eq, Nat.cast_one] + rw [← Nat.cast_le (α := ℕ∞), (toFinite _).cast_ncard_eq, Nat.cast_one] apply encard_singleton_inter #align set.ncard_singleton_inter Set.ncard_singleton_inter section InsertErase @[simp] theorem ncard_insert_of_not_mem (h : a ∉ s) (hs : s.Finite := by toFinite_tac) : (insert a s).ncard = s.ncard + 1 := by - rw [←Nat.cast_inj (R := ℕ∞), (hs.insert a).cast_ncard_eq, Nat.cast_add, Nat.cast_one, + rw [← Nat.cast_inj (R := ℕ∞), (hs.insert a).cast_ncard_eq, Nat.cast_add, Nat.cast_one, hs.cast_ncard_eq, encard_insert_of_not_mem h] #align set.ncard_insert_of_not_mem Set.ncard_insert_of_not_mem @@ -653,7 +653,7 @@ theorem ncard_image_of_injOn (H : Set.InjOn f s) : (f '' s).ncard = s.ncard := theorem injOn_of_ncard_image_eq (h : (f '' s).ncard = s.ncard) (hs : s.Finite := by toFinite_tac) : Set.InjOn f s := by - rw [←Nat.cast_inj (R := ℕ∞), hs.cast_ncard_eq, (hs.image _).cast_ncard_eq] at h + rw [← Nat.cast_inj (R := ℕ∞), hs.cast_ncard_eq, (hs.image _).cast_ncard_eq] at h exact hs.injOn_of_encard_image_eq h #align set.inj_on_of_ncard_image_eq Set.injOn_of_ncard_image_eq @@ -688,7 +688,7 @@ theorem fiber_ncard_ne_zero_iff_mem_image {y : β} (hs : s.Finite := by toFinite { x : Subtype P | (x : α) ∈ s }.ncard = (s ∩ setOf P).ncard := by convert (ncard_image_of_injective _ (@Subtype.coe_injective _ P)).symm ext x - simp [←and_assoc, exists_eq_right] + simp [← and_assoc, exists_eq_right] #align set.ncard_subtype Set.ncard_subtype theorem ncard_inter_le_ncard_left (s t : Set α) (hs : s.Finite := by toFinite_tac) : @@ -704,7 +704,7 @@ theorem ncard_inter_le_ncard_right (s t : Set α) (ht : t.Finite := by toFinite_ theorem eq_of_subset_of_ncard_le (h : s ⊆ t) (h' : t.ncard ≤ s.ncard) (ht : t.Finite := by toFinite_tac) : s = t := ht.eq_of_subset_of_encard_le h - (by rwa [←Nat.cast_le (α := ℕ∞), ht.cast_ncard_eq, (ht.subset h).cast_ncard_eq] at h') + (by rwa [← Nat.cast_le (α := ℕ∞), ht.cast_ncard_eq, (ht.subset h).cast_ncard_eq] at h') #align set.eq_of_subset_of_ncard_le Set.eq_of_subset_of_ncard_le theorem subset_iff_eq_of_ncard_le (h : t.ncard ≤ s.ncard) (ht : t.Finite := by toFinite_tac) : @@ -724,7 +724,7 @@ theorem sep_of_ncard_eq {P : α → Prop} (h : { x ∈ s | P x }.ncard = s.ncard theorem ncard_lt_ncard (h : s ⊂ t) (ht : t.Finite := by toFinite_tac) : s.ncard < t.ncard := by - rw [←Nat.cast_lt (α := ℕ∞), ht.cast_ncard_eq, (ht.subset h.subset).cast_ncard_eq] + rw [← Nat.cast_lt (α := ℕ∞), ht.cast_ncard_eq, (ht.subset h.subset).cast_ncard_eq] exact ht.encard_lt_encard h #align set.ncard_lt_ncard Set.ncard_lt_ncard @@ -755,7 +755,7 @@ theorem ncard_congr {t : Set β} (f : ∀ a ∈ s, β) (h₁ : ∀ a ha, f a ha obtain ⟨a, ha, rfl⟩ := h₃ y hy simp only [Subtype.mk.injEq, Subtype.exists] exact ⟨_, ha, rfl⟩ - simp_rw [←Nat.card_coe_set_eq] + simp_rw [← Nat.card_coe_set_eq] exact Nat.card_congr (Equiv.ofBijective f' hbij) #align set.ncard_congr Set.ncard_congr @@ -800,7 +800,7 @@ theorem surj_on_of_inj_on_of_ncard_le {t : Set β} (f : ∀ a ∈ s, β) (hf : · intros a₁ a₂ ha₁ ha₂ h rw [mem_toFinset] at ha₁ ha₂ exact hinj _ _ ha₁ ha₂ h - rwa [←ncard_eq_toFinset_card', ←ncard_eq_toFinset_card'] + rwa [← ncard_eq_toFinset_card', ← ncard_eq_toFinset_card'] #align set.surj_on_of_inj_on_of_ncard_le Set.surj_on_of_inj_on_of_ncard_le theorem inj_on_of_surj_on_of_ncard_le {t : Set β} (f : ∀ a ∈ s, β) (hf : ∀ a ha, f a ha ∈ t) @@ -820,7 +820,7 @@ theorem inj_on_of_surj_on_of_ncard_le {t : Set β} (f : ∀ a ∈ s, β) (hf : exact @Finset.inj_on_of_surj_on_of_card_le _ _ _ t.toFinset f'' (fun a ha ↦ by { rw [mem_toFinset] at ha ⊢; exact hf a ha }) (by simpa) - (by { rwa [←ncard_eq_toFinset_card', ←ncard_eq_toFinset_card'] }) a₁ a₂ + (by { rwa [← ncard_eq_toFinset_card', ← ncard_eq_toFinset_card'] }) a₁ a₂ (by simpa) (by simpa) (by simpa) #align set.inj_on_of_surj_on_of_ncard_le Set.inj_on_of_surj_on_of_ncard_le @@ -883,7 +883,7 @@ theorem le_ncard_diff (s t : Set α) (hs : s.Finite := by toFinite_tac) : theorem ncard_diff_add_ncard (s t : Set α) (hs : s.Finite := by toFinite_tac) (ht : t.Finite := by toFinite_tac) : (s \ t).ncard + t.ncard = (s ∪ t).ncard := by - rw [←ncard_union_eq disjoint_sdiff_left (hs.diff _) ht, diff_union_self] + rw [← ncard_union_eq disjoint_sdiff_left (hs.diff _) ht, diff_union_self] #align set.ncard_diff_add_ncard Set.ncard_diff_add_ncard theorem diff_nonempty_of_ncard_lt_ncard (h : s.ncard < t.ncard) (hs : s.Finite := by toFinite_tac) : @@ -899,7 +899,7 @@ theorem exists_mem_not_mem_of_ncard_lt_ncard (h : s.ncard < t.ncard) @[simp] theorem ncard_inter_add_ncard_diff_eq_ncard (s t : Set α) (hs : s.Finite := by toFinite_tac) : (s ∩ t).ncard + (s \ t).ncard = s.ncard := by - rw [←ncard_union_eq (disjoint_of_subset_left (inter_subset_right _ _) disjoint_sdiff_right) + rw [← ncard_union_eq (disjoint_of_subset_left (inter_subset_right _ _) disjoint_sdiff_right) (hs.inter_of_left _) (hs.diff _), union_comm, diff_union_inter] #align set.ncard_inter_add_ncard_diff_eq_ncard Set.ncard_inter_add_ncard_diff_eq_ncard @@ -1004,7 +1004,7 @@ theorem exists_eq_insert_iff_ncard (hs : s.Finite := by toFinite_tac) : classical cases' t.finite_or_infinite with ht ht · rw [ncard_eq_toFinset_card _ hs, ncard_eq_toFinset_card _ ht, - ←@Finite.toFinset_subset_toFinset _ _ _ hs ht, ←Finset.exists_eq_insert_iff] + ← @Finite.toFinset_subset_toFinset _ _ _ hs ht, ← Finset.exists_eq_insert_iff] convert Iff.rfl using 2; simp ext x simp [Finset.ext_iff, Set.ext_iff] @@ -1097,14 +1097,14 @@ theorem ncard_eq_succ {n : ℕ} (hs : s.Finite := by toFinite_tac) : #align set.ncard_eq_succ Set.ncard_eq_succ theorem ncard_eq_two : s.ncard = 2 ↔ ∃ x y, x ≠ y ∧ s = {x, y} := by - rw [←encard_eq_two, ncard_def, ←Nat.cast_inj (R := ℕ∞), Nat.cast_ofNat] + rw [← encard_eq_two, ncard_def, ← Nat.cast_inj (R := ℕ∞), Nat.cast_ofNat] refine' ⟨fun h ↦ _, fun h ↦ _⟩ · rwa [ENat.coe_toNat] at h; rintro h'; simp [h'] at h rw [h]; rfl #align set.ncard_eq_two Set.ncard_eq_two theorem ncard_eq_three : s.ncard = 3 ↔ ∃ x y z, x ≠ y ∧ x ≠ z ∧ y ≠ z ∧ s = {x, y, z} := by - rw [←encard_eq_three, ncard_def, ←Nat.cast_inj (R := ℕ∞), Nat.cast_ofNat] + rw [← encard_eq_three, ncard_def, ← Nat.cast_inj (R := ℕ∞), Nat.cast_ofNat] refine' ⟨fun h ↦ _, fun h ↦ _⟩ · rwa [ENat.coe_toNat] at h; rintro h'; simp [h'] at h rw [h]; rfl diff --git a/Mathlib/Data/Set/Finite.lean b/Mathlib/Data/Set/Finite.lean index ea186a579271d..f883762ff0cb6 100644 --- a/Mathlib/Data/Set/Finite.lean +++ b/Mathlib/Data/Set/Finite.lean @@ -964,7 +964,7 @@ protected theorem infinite_prod : (s ×ˢ t).Infinite ↔ s.Infinite ∧ t.Nonempty ∨ t.Infinite ∧ s.Nonempty := by refine' ⟨fun h => _, _⟩ · simp_rw [Set.Infinite, @and_comm ¬_, ← not_imp] - by_contra' + by_contra! exact h ((this.1 h.nonempty.snd).prod $ this.2 h.nonempty.fst) · rintro (h | h) · exact h.1.prod_left h.2 @@ -1312,7 +1312,7 @@ alias ⟨_, Infinite.to_subtype⟩ := infinite_coe_iff #align set.infinite.to_subtype Set.Infinite.to_subtype lemma Infinite.exists_not_mem_finite (hs : s.Infinite) (ht : t.Finite) : ∃ a, a ∈ s ∧ a ∉ t := by - by_contra' h; exact hs $ ht.subset h + by_contra! h; exact hs $ ht.subset h lemma Infinite.exists_not_mem_finset (hs : s.Infinite) (t : Finset α) : ∃ a ∈ s, a ∉ t := hs.exists_not_mem_finite t.finite_toSet @@ -1322,7 +1322,7 @@ section Infinite variable [Infinite α] lemma Finite.exists_not_mem (hs : s.Finite) : ∃ a, a ∉ s := by - by_contra' h; exact infinite_univ (hs.subset fun a _ ↦ h _) + by_contra! h; exact infinite_univ (hs.subset fun a _ ↦ h _) lemma _root_.Finset.exists_not_mem (s : Finset α) : ∃ a, a ∉ s := s.finite_toSet.exists_not_mem diff --git a/Mathlib/Data/Set/Function.lean b/Mathlib/Data/Set/Function.lean index 8d976c1b5e019..a35d1539416cf 100644 --- a/Mathlib/Data/Set/Function.lean +++ b/Mathlib/Data/Set/Function.lean @@ -1302,7 +1302,7 @@ theorem injOn_iff_invFunOn_image_image_eq_self [Nonempty α] : theorem _root_.Function.invFunOn_injOn_image [Nonempty α] (f : α → β) (s : Set α) : Set.InjOn (invFunOn f s) (f '' s) := by rintro _ ⟨x, hx, rfl⟩ _ ⟨x', hx', rfl⟩ he - rw [←invFunOn_apply_eq (f := f) hx, he, invFunOn_apply_eq (f := f) hx'] + rw [← invFunOn_apply_eq (f := f) hx, he, invFunOn_apply_eq (f := f) hx'] theorem _root_.Function.invFunOn_image_image_subset [Nonempty α] (f : α → β) (s : Set α) : (invFunOn f s) '' (f '' s) ⊆ s := by diff --git a/Mathlib/Data/Set/Image.lean b/Mathlib/Data/Set/Image.lean index c7d8f114341f3..2c516ce796bc6 100644 --- a/Mathlib/Data/Set/Image.lean +++ b/Mathlib/Data/Set/Image.lean @@ -436,7 +436,7 @@ theorem mem_image_iff_of_inverse {f : α → β} {g : β → α} {b : β} {s : S #align set.mem_image_iff_of_inverse Set.mem_image_iff_of_inverse theorem image_compl_subset {f : α → β} {s : Set α} (H : Injective f) : f '' sᶜ ⊆ (f '' s)ᶜ := - Disjoint.subset_compl_left <| by simp [disjoint_iff_inf_le, ←image_inter H] + Disjoint.subset_compl_left <| by simp [disjoint_iff_inf_le, ← image_inter H] #align set.image_compl_subset Set.image_compl_subset theorem subset_image_compl {f : α → β} {s : Set α} (H : Surjective f) : (f '' s)ᶜ ⊆ f '' sᶜ := diff --git a/Mathlib/Data/Set/Intervals/Basic.lean b/Mathlib/Data/Set/Intervals/Basic.lean index 896c3b443fd55..75b44906da590 100644 --- a/Mathlib/Data/Set/Intervals/Basic.lean +++ b/Mathlib/Data/Set/Intervals/Basic.lean @@ -1180,8 +1180,7 @@ theorem Ico_eq_Ico_iff (h : a₁ < b₁ ∨ a₂ < b₂) : Ico a₁ b₁ = Ico a lemma Ici_eq_singleton_iff_isTop {x : α} : (Ici x = {x}) ↔ IsTop x := by refine ⟨fun h y ↦ ?_, fun h ↦ by ext y; simp [(h y).ge_iff_eq]⟩ - by_contra H - push_neg at H + by_contra! H have : y ∈ Ici x := H.le rw [h, mem_singleton_iff] at this exact lt_irrefl y (this.le.trans_lt H) diff --git a/Mathlib/Data/Set/Intervals/UnorderedInterval.lean b/Mathlib/Data/Set/Intervals/UnorderedInterval.lean index 9637bdb8a64a6..23e6724a4a0fb 100644 --- a/Mathlib/Data/Set/Intervals/UnorderedInterval.lean +++ b/Mathlib/Data/Set/Intervals/UnorderedInterval.lean @@ -187,12 +187,12 @@ variable [Lattice β] {f : α → β} {s : Set α} {a b : α} lemma _root_.MonotoneOn.mapsTo_uIcc (hf : MonotoneOn f (uIcc a b)) : MapsTo f (uIcc a b) (uIcc (f a) (f b)) := by - rw [uIcc, uIcc, ←hf.map_sup, ←hf.map_inf] <;> + rw [uIcc, uIcc, ← hf.map_sup, ← hf.map_inf] <;> apply_rules [left_mem_uIcc, right_mem_uIcc, hf.mapsTo_Icc] lemma _root_.AntitoneOn.mapsTo_uIcc (hf : AntitoneOn f (uIcc a b)) : MapsTo f (uIcc a b) (uIcc (f a) (f b)) := by - rw [uIcc, uIcc, ←hf.map_sup, ←hf.map_inf] <;> + rw [uIcc, uIcc, ← hf.map_sup, ← hf.map_inf] <;> apply_rules [left_mem_uIcc, right_mem_uIcc, hf.mapsTo_Icc] lemma _root_.Monotone.mapsTo_uIcc (hf : Monotone f) : MapsTo f (uIcc a b) (uIcc (f a) (f b)) := diff --git a/Mathlib/Data/Set/NAry.lean b/Mathlib/Data/Set/NAry.lean index f5761329cf7ff..c60dd9b7deaee 100644 --- a/Mathlib/Data/Set/NAry.lean +++ b/Mathlib/Data/Set/NAry.lean @@ -111,7 +111,7 @@ lemma image_prod : (fun x : α × β ↦ f x.1 x.2) '' s ×ˢ t = image2 f s t : -- Porting note: Removing `simp` - LHS does not simplify lemma image2_curry (f : α × β → γ) (s : Set α) (t : Set β) : image2 (fun a b ↦ f (a, b)) s t = f '' s ×ˢ t := by - simp [←image_uncurry_prod, uncurry] + simp [← image_uncurry_prod, uncurry] #align set.image2_curry Set.image2_curry theorem image2_swap (s : Set α) (t : Set β) : image2 f s t = image2 (fun a b => f b a) t s := by @@ -135,12 +135,12 @@ theorem image2_union_right : image2 f s (t ∪ t') = image2 f s t ∪ image2 f s lemma image2_inter_left (hf : Injective2 f) : image2 f (s ∩ s') t = image2 f s t ∩ image2 f s' t := by - simp_rw [←image_uncurry_prod, inter_prod, image_inter hf.uncurry] + simp_rw [← image_uncurry_prod, inter_prod, image_inter hf.uncurry] #align set.image2_inter_left Set.image2_inter_left lemma image2_inter_right (hf : Injective2 f) : image2 f s (t ∩ t') = image2 f s t ∩ image2 f s t' := by - simp_rw [←image_uncurry_prod, prod_inter, image_inter hf.uncurry] + simp_rw [← image_uncurry_prod, prod_inter, image_inter hf.uncurry] #align set.image2_inter_right Set.image2_inter_right @[simp] diff --git a/Mathlib/Data/Set/Pointwise/SMul.lean b/Mathlib/Data/Set/Pointwise/SMul.lean index fc7cd705c3915..4ec93b98e086f 100644 --- a/Mathlib/Data/Set/Pointwise/SMul.lean +++ b/Mathlib/Data/Set/Pointwise/SMul.lean @@ -563,7 +563,7 @@ scoped[Pointwise] attribute [instance] Set.distribMulActionSet Set.mulDistribMul instance [Zero α] [Zero β] [SMul α β] [NoZeroSMulDivisors α β] : NoZeroSMulDivisors (Set α) (Set β) := ⟨fun {s t} h ↦ by - by_contra' H + by_contra! H have hst : (s • t).Nonempty := h.symm.subst zero_nonempty rw [Ne.def, ← hst.of_smul_left.subset_zero_iff, Ne.def, ← hst.of_smul_right.subset_zero_iff] at H @@ -574,7 +574,7 @@ instance [Zero α] [Zero β] [SMul α β] [NoZeroSMulDivisors α β] : instance noZeroSMulDivisors_set [Zero α] [Zero β] [SMul α β] [NoZeroSMulDivisors α β] : NoZeroSMulDivisors α (Set β) := ⟨fun {a s} h ↦ by - by_contra' H + by_contra! H have hst : (a • s).Nonempty := h.symm.subst zero_nonempty rw [Ne.def, Ne.def, ← hst.of_image.subset_zero_iff, not_subset] at H obtain ⟨ha, b, ht, hb⟩ := H diff --git a/Mathlib/Data/Setoid/Partition.lean b/Mathlib/Data/Setoid/Partition.lean index 5b35084703b52..144f049dd05ac 100644 --- a/Mathlib/Data/Setoid/Partition.lean +++ b/Mathlib/Data/Setoid/Partition.lean @@ -213,7 +213,7 @@ noncomputable def quotientEquivClasses (r : Setoid α) : Quotient r ≃ Setoid.c simp only [Subtype.ext_iff, Quotient.lift_mk, Subtype.ext_iff] at h_eq apply Quotient.sound show a ∈ { x | Setoid.r x b } - rw [←h_eq] + rw [← h_eq] exact Setoid.refl a · rw [Quot.surjective_lift] intro ⟨c, a, hc⟩ diff --git a/Mathlib/Data/Sign.lean b/Mathlib/Data/Sign.lean index ae98a5158f3eb..e4f20d123046f 100644 --- a/Mathlib/Data/Sign.lean +++ b/Mathlib/Data/Sign.lean @@ -360,7 +360,7 @@ theorem sign_ne_zero : sign a ≠ 0 ↔ a ≠ 0 := theorem sign_nonneg_iff : 0 ≤ sign a ↔ 0 ≤ a := by rcases lt_trichotomy 0 a with (h | h | h) · simp [h, h.le] - · simp [←h] + · simp [← h] · simp [h, h.not_le] #align sign_nonneg_iff sign_nonneg_iff @@ -368,7 +368,7 @@ theorem sign_nonneg_iff : 0 ≤ sign a ↔ 0 ≤ a := by theorem sign_nonpos_iff : sign a ≤ 0 ↔ a ≤ 0 := by rcases lt_trichotomy 0 a with (h | h | h) · simp [h, h.not_le] - · simp [←h] + · simp [← h] · simp [h, h.le] #align sign_nonpos_iff sign_nonpos_iff diff --git a/Mathlib/Data/Sum/Lattice.lean b/Mathlib/Data/Sum/Lattice.lean index 24f04370881c8..316e8b67f10a4 100644 --- a/Mathlib/Data/Sum/Lattice.lean +++ b/Mathlib/Data/Sum/Lattice.lean @@ -107,7 +107,7 @@ instance instDistribLattice [DistribLattice α] [DistribLattice β] : DistribLat le_sup_inf := by simp only [Lex.forall, Sum.forall, ge_iff_le, inl_le_inl_iff, inr_le_inr_iff, sup_le_iff, le_sup_left, true_and, inl_le_inr, not_inr_le_inl, le_inf_iff, sup_of_le_right, and_self, - inf_of_le_left, le_refl, implies_true, and_true, inf_of_le_right, sup_of_le_left, ←inl_sup, - ←inr_sup, ←inl_inf, ←inr_inf, sup_inf_left, le_rfl] + inf_of_le_left, le_refl, implies_true, and_true, inf_of_le_right, sup_of_le_left, ← inl_sup, + ← inr_sup, ← inl_inf, ← inr_inf, sup_inf_left, le_rfl] end Sum.Lex diff --git a/Mathlib/Data/Sym/Sym2.lean b/Mathlib/Data/Sym/Sym2.lean index d41efad7d50e2..cb6513f988586 100644 --- a/Mathlib/Data/Sym/Sym2.lean +++ b/Mathlib/Data/Sym/Sym2.lean @@ -571,7 +571,7 @@ private def fromVector : Vector α 2 → α × α private theorem perm_card_two_iff {a₁ b₁ a₂ b₂ : α} : [a₁, b₁].Perm [a₂, b₂] ↔ a₁ = a₂ ∧ b₁ = b₂ ∨ a₁ = b₂ ∧ b₁ = a₂ := { mp := by - simp only [←Multiset.coe_eq_coe, ←Multiset.cons_coe, Multiset.coe_nil, Multiset.cons_zero, + simp only [← Multiset.coe_eq_coe, ← Multiset.cons_coe, Multiset.coe_nil, Multiset.cons_zero, Multiset.cons_eq_cons, Multiset.singleton_inj, ne_eq, Multiset.singleton_eq_cons_iff, exists_eq_right_right, and_true] tauto diff --git a/Mathlib/Data/Vector/Basic.lean b/Mathlib/Data/Vector/Basic.lean index 9bfd215bf638c..e3001bb438f2f 100644 --- a/Mathlib/Data/Vector/Basic.lean +++ b/Mathlib/Data/Vector/Basic.lean @@ -169,7 +169,7 @@ theorem get_tail (x : Vector α n) (i) : cases' i with i ih; dsimp rcases x with ⟨_ | _, h⟩ <;> try rfl rw [List.length] at h - rw [←h] at ih + rw [← h] at ih contradiction #align vector.nth_tail Vector.get_tail diff --git a/Mathlib/Data/ZMod/Basic.lean b/Mathlib/Data/ZMod/Basic.lean index b580604a670f5..2af38868838e5 100644 --- a/Mathlib/Data/ZMod/Basic.lean +++ b/Mathlib/Data/ZMod/Basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes -/ import Mathlib.Algebra.CharP.Basic +import Mathlib.RingTheory.Ideal.Operations import Mathlib.Data.Fintype.Units import Mathlib.Data.Nat.Parity import Mathlib.Tactic.FinCases @@ -842,7 +843,7 @@ instance subsingleton_units : Subsingleton (ZMod 2)ˣ := theorem add_self_eq_zero_iff_eq_zero {n : ℕ} (hn : Odd n) {a : ZMod n} : a + a = 0 ↔ a = 0 := by rw [Nat.odd_iff, ← Nat.two_dvd_ne_zero, ← Nat.prime_two.coprime_iff_not_dvd] at hn - rw [←mul_two, ←@Nat.cast_two (ZMod n), ←ZMod.coe_unitOfCoprime 2 hn, Units.mul_left_eq_zero] + rw [← mul_two, ← @Nat.cast_two (ZMod n), ← ZMod.coe_unitOfCoprime 2 hn, Units.mul_left_eq_zero] theorem ne_neg_self {n : ℕ} (hn : Odd n) {a : ZMod n} (ha : a ≠ 0) : a ≠ -a := by rwa [Ne, eq_neg_iff_add_eq_zero, add_self_eq_zero_iff_eq_zero hn] diff --git a/Mathlib/Data/ZMod/IntUnitsPower.lean b/Mathlib/Data/ZMod/IntUnitsPower.lean index 6aa8079faa15a..0e39556b0dedf 100644 --- a/Mathlib/Data/ZMod/IntUnitsPower.lean +++ b/Mathlib/Data/ZMod/IntUnitsPower.lean @@ -39,12 +39,12 @@ instance : Module (ZMod 2) (Additive ℤˣ) where one_smul au := Additive.toMul.injective <| pow_one _ mul_smul z₁ z₂ au := Additive.toMul.injective <| by dsimp only [ZMod.smul_units_def, toMul_nsmul] - rw [←pow_mul, ZMod.val_mul, ←Int.units_pow_eq_pow_mod_two, mul_comm] + rw [← pow_mul, ZMod.val_mul, ← Int.units_pow_eq_pow_mod_two, mul_comm] smul_zero z := Additive.toMul.injective <| one_pow _ smul_add z au₁ au₂ := Additive.toMul.injective <| mul_pow _ _ _ add_smul z₁ z₂ au := Additive.toMul.injective <| by dsimp only [ZMod.smul_units_def, toMul_nsmul, toMul_add] - rw [←pow_add, ZMod.val_add, ←Int.units_pow_eq_pow_mod_two] + rw [← pow_add, ZMod.val_add, ← Int.units_pow_eq_pow_mod_two] zero_smul au := Additive.toMul.injective <| pow_zero (Additive.toMul au) section CommSemiring @@ -69,7 +69,7 @@ example : Int.instUnitsPow = DivInvMonoid.Pow := rfl @[norm_cast] lemma uzpow_natCast (u : ℤˣ) (n : ℕ) : u ^ (n : R) = u ^ n := by change Additive.toMul ((n : R) • Additive.ofMul u) = _ - rw [←nsmul_eq_smul_cast, toMul_nsmul, toMul_ofMul] + rw [← nsmul_eq_smul_cast, toMul_nsmul, toMul_ofMul] -- See note [no_index around OfNat.ofNat] lemma uzpow_ofNat (s : ℤˣ) (n : ℕ) [n.AtLeastTwo] : @@ -107,6 +107,6 @@ lemma uzpow_neg (s : ℤˣ) (x : R) : s ^ (-x) = (s ^ x)⁻¹ := @[norm_cast] lemma uzpow_intCast (u : ℤˣ) (z : ℤ) : u ^ (z : R) = u ^ z := by change Additive.toMul ((z : R) • Additive.ofMul u) = _ - rw [←zsmul_eq_smul_cast, toMul_zsmul, toMul_ofMul] + rw [← zsmul_eq_smul_cast, toMul_zsmul, toMul_ofMul] end CommRing diff --git a/Mathlib/Dynamics/Circle/RotationNumber/TranslationNumber.lean b/Mathlib/Dynamics/Circle/RotationNumber/TranslationNumber.lean index 320e0c9ba1773..58a3f2ce6eb31 100644 --- a/Mathlib/Dynamics/Circle/RotationNumber/TranslationNumber.lean +++ b/Mathlib/Dynamics/Circle/RotationNumber/TranslationNumber.lean @@ -927,10 +927,10 @@ theorem lt_translationNumber_of_forall_add_lt (hf : Continuous f) {z : ℝ} (hz such that `f x = x + τ f`. -/ theorem exists_eq_add_translationNumber (hf : Continuous f) : ∃ x, f x = x + τ f := by obtain ⟨a, ha⟩ : ∃ x, f x ≤ x + τ f := by - by_contra' H + by_contra! H exact lt_irrefl _ (f.lt_translationNumber_of_forall_add_lt hf H) obtain ⟨b, hb⟩ : ∃ x, x + τ f ≤ f x := by - by_contra' H + by_contra! H exact lt_irrefl _ (f.translationNumber_lt_of_forall_lt_add hf H) exact intermediate_value_univ₂ hf (continuous_id.add continuous_const) ha hb #align circle_deg1_lift.exists_eq_add_translation_number CircleDeg1Lift.exists_eq_add_translationNumber diff --git a/Mathlib/Dynamics/Ergodic/Conservative.lean b/Mathlib/Dynamics/Ergodic/Conservative.lean index e839ba629d85d..4cf47e8a1cce5 100644 --- a/Mathlib/Dynamics/Ergodic/Conservative.lean +++ b/Mathlib/Dynamics/Ergodic/Conservative.lean @@ -54,14 +54,16 @@ open Measure *conservative* if for any measurable set `s` of positive measure there exists `x ∈ s` such that `x` returns back to `s` under some iteration of `f`. -/ structure Conservative (f : α → α) (μ : Measure α) extends QuasiMeasurePreserving f μ μ : Prop where - exists_mem_image_mem : - ∀ ⦃s⦄, MeasurableSet s → μ s ≠ 0 → ∃ x ∈ s, ∃ (m : _) (_ : m ≠ 0), f^[m] x ∈ s + /-- If `f` is a conservative self-map and `s` is a measurable set of nonzero measure, + then there exists a point `x ∈ s` that returns to `s` under a non-zero iteration of `f`. -/ + exists_mem_iterate_mem : + ∀ ⦃s⦄, MeasurableSet s → μ s ≠ 0 → ∃ x ∈ s, ∃ m, m ≠ 0 ∧ f^[m] x ∈ s #align measure_theory.conservative MeasureTheory.Conservative /-- A self-map preserving a finite measure is conservative. -/ protected theorem MeasurePreserving.conservative [IsFiniteMeasure μ] (h : MeasurePreserving f μ μ) : Conservative f μ := - ⟨h.quasiMeasurePreserving, fun _ hsm h0 => h.exists_mem_image_mem hsm h0⟩ + ⟨h.quasiMeasurePreserving, fun _ hsm h0 => h.exists_mem_iterate_mem hsm h0⟩ #align measure_theory.measure_preserving.conservative MeasureTheory.MeasurePreserving.conservative namespace Conservative @@ -69,7 +71,7 @@ namespace Conservative /-- The identity map is conservative w.r.t. any measure. -/ protected theorem id (μ : Measure α) : Conservative id μ := { toQuasiMeasurePreserving := QuasiMeasurePreserving.id μ - exists_mem_image_mem := fun _ _ h0 => + exists_mem_iterate_mem := fun _ _ h0 => let ⟨x, hx⟩ := nonempty_of_measure_ne_zero h0 ⟨x, hx, 1, one_ne_zero, hx⟩ } #align measure_theory.conservative.id MeasureTheory.Conservative.id @@ -96,7 +98,7 @@ theorem frequently_measure_inter_ne_zero (hf : Conservative f μ) (hs : Measurab rw [← inter_iUnion₂] rfl have : μ ((s ∩ f^[n] ⁻¹' s) \ T) ≠ 0 := by rwa [measure_diff_null hμT] - rcases hf.exists_mem_image_mem ((hs.inter (hf.measurable.iterate n hs)).diff hT) this with + rcases hf.exists_mem_iterate_mem ((hs.inter (hf.measurable.iterate n hs)).diff hT) this with ⟨x, ⟨⟨hxs, _⟩, hxT⟩, m, hm0, ⟨_, hxm⟩, _⟩ refine' hxT ⟨hxs, mem_iUnion₂.2 ⟨n + m, _, _⟩⟩ · exact add_le_add hn (Nat.one_le_of_lt <| pos_iff_ne_zero.2 hm0) diff --git a/Mathlib/Dynamics/Ergodic/MeasurePreserving.lean b/Mathlib/Dynamics/Ergodic/MeasurePreserving.lean index 082ea7de7c434..0067365ee87f1 100644 --- a/Mathlib/Dynamics/Ergodic/MeasurePreserving.lean +++ b/Mathlib/Dynamics/Ergodic/MeasurePreserving.lean @@ -34,6 +34,7 @@ variable {α β γ δ : Type*} [MeasurableSpace α] [MeasurableSpace β] [Measur namespace MeasureTheory open Measure Function Set +open scoped BigOperators variable {μa : Measure α} {μb : Measure β} {μc : Measure γ} {μd : Measure δ} @@ -159,32 +160,31 @@ lemma measure_symmDiff_preimage_iterate_le /-- If `μ univ < n * μ s` and `f` is a map preserving measure `μ`, then for some `x ∈ s` and `0 < m < n`, `f^[m] x ∈ s`. -/ -theorem exists_mem_image_mem_of_volume_lt_mul_volume (hf : MeasurePreserving f μ μ) +theorem exists_mem_iterate_mem_of_volume_lt_mul_volume (hf : MeasurePreserving f μ μ) (hs : MeasurableSet s) {n : ℕ} (hvol : μ (Set.univ : Set α) < n * μ s) : ∃ x ∈ s, ∃ m ∈ Set.Ioo 0 n, f^[m] x ∈ s := by - have A : ∀ m, MeasurableSet (f^[m] ⁻¹' s) := fun m => (hf.iterate m).measurable hs - have B : ∀ m, μ (f^[m] ⁻¹' s) = μ s := fun m => (hf.iterate m).measure_preimage hs - have : μ (Set.univ : Set α) < (Finset.range n).sum fun m => μ (f^[m] ⁻¹' s) := by - simpa only [B, nsmul_eq_mul, Finset.sum_const, Finset.card_range] - rcases exists_nonempty_inter_of_measure_univ_lt_sum_measure μ (fun m _ => A m) this with - ⟨i, hi, j, hj, hij, x, hxi, hxj⟩ + have A : ∀ m, MeasurableSet (f^[m] ⁻¹' s) := fun m ↦ (hf.iterate m).measurable hs + have B : ∀ m, μ (f^[m] ⁻¹' s) = μ s := fun m ↦ (hf.iterate m).measure_preimage hs + have : μ (univ : Set α) < ∑ m in Finset.range n, μ (f^[m] ⁻¹' s) := by simpa [B] + obtain ⟨i, hi, j, hj, hij, x, hxi : f^[i] x ∈ s, hxj : f^[j] x ∈ s⟩ : + ∃ i < n, ∃ j < n, i ≠ j ∧ (f^[i] ⁻¹' s ∩ f^[j] ⁻¹' s).Nonempty := by + simpa using exists_nonempty_inter_of_measure_univ_lt_sum_measure μ (fun m _ ↦ A m) this wlog hlt : i < j generalizing i j · exact this j hj i hi hij.symm hxj hxi (hij.lt_or_lt.resolve_left hlt) - simp only [Set.mem_preimage, Finset.mem_range] at hi hj hxi hxj - refine' ⟨f^[i] x, hxi, j - i, ⟨tsub_pos_of_lt hlt, lt_of_le_of_lt (j.sub_le i) hj⟩, _⟩ + refine ⟨f^[i] x, hxi, j - i, ⟨tsub_pos_of_lt hlt, lt_of_le_of_lt (j.sub_le i) hj⟩, ?_⟩ rwa [← iterate_add_apply, tsub_add_cancel_of_le hlt.le] -#align measure_theory.measure_preserving.exists_mem_image_mem_of_volume_lt_mul_volume MeasureTheory.MeasurePreserving.exists_mem_image_mem_of_volume_lt_mul_volume +#align measure_theory.measure_preserving.exists_mem_image_mem_of_volume_lt_mul_volume MeasureTheory.MeasurePreserving.exists_mem_iterate_mem_of_volume_lt_mul_volume /-- A self-map preserving a finite measure is conservative: if `μ s ≠ 0`, then at least one point `x ∈ s` comes back to `s` under iterations of `f`. Actually, a.e. point of `s` comes back to `s` infinitely many times, see `MeasureTheory.MeasurePreserving.conservative` and theorems about `MeasureTheory.Conservative`. -/ -theorem exists_mem_image_mem [IsFiniteMeasure μ] (hf : MeasurePreserving f μ μ) - (hs : MeasurableSet s) (hs' : μ s ≠ 0) : ∃ x ∈ s, ∃ (m : _) (_ : m ≠ 0), f^[m] x ∈ s := by +theorem exists_mem_iterate_mem [IsFiniteMeasure μ] (hf : MeasurePreserving f μ μ) + (hs : MeasurableSet s) (hs' : μ s ≠ 0) : ∃ x ∈ s, ∃ m, m ≠ 0 ∧ f^[m] x ∈ s := by rcases ENNReal.exists_nat_mul_gt hs' (measure_ne_top μ (Set.univ : Set α)) with ⟨N, hN⟩ - rcases hf.exists_mem_image_mem_of_volume_lt_mul_volume hs hN with ⟨x, hx, m, hm, hmx⟩ + rcases hf.exists_mem_iterate_mem_of_volume_lt_mul_volume hs hN with ⟨x, hx, m, hm, hmx⟩ exact ⟨x, hx, m, hm.1.ne', hmx⟩ -#align measure_theory.measure_preserving.exists_mem_image_mem MeasureTheory.MeasurePreserving.exists_mem_image_mem +#align measure_theory.measure_preserving.exists_mem_image_mem MeasureTheory.MeasurePreserving.exists_mem_iterate_mem end MeasurePreserving diff --git a/Mathlib/Dynamics/Flow.lean b/Mathlib/Dynamics/Flow.lean index ff109a643da99..d3c88ee09b447 100644 --- a/Mathlib/Dynamics/Flow.lean +++ b/Mathlib/Dynamics/Flow.lean @@ -178,7 +178,7 @@ def reverse : Flow τ α where -- Porting note: Homeomorphism.continuous_invFun : Continuous invFun := by continuity @[continuity] theorem continuous_toFun (t : τ) : Continuous (ϕ.toFun t) := by - rw [←curry_uncurry ϕ.toFun] + rw [← curry_uncurry ϕ.toFun] apply continuous_curry exact ϕ.cont' diff --git a/Mathlib/Dynamics/PeriodicPts.lean b/Mathlib/Dynamics/PeriodicPts.lean index de8133e6a7144..b559b69559931 100644 --- a/Mathlib/Dynamics/PeriodicPts.lean +++ b/Mathlib/Dynamics/PeriodicPts.lean @@ -343,7 +343,7 @@ theorem minimalPeriod_apply (hx : x ∈ periodicPts f) : minimalPeriod f (f x) = theorem le_of_lt_minimalPeriod_of_iterate_eq {m n : ℕ} (hm : m < minimalPeriod f x) (hmn : f^[m] x = f^[n] x) : m ≤ n := by - by_contra' hmn' + by_contra! hmn' rw [← Nat.add_sub_of_le hmn'.le, add_comm, iterate_add_apply] at hmn exact ((IsPeriodicPt.minimalPeriod_le (tsub_pos_of_lt hmn') diff --git a/Mathlib/FieldTheory/Adjoin.lean b/Mathlib/FieldTheory/Adjoin.lean index a2d8e81091f5f..ceb68f0d6d7ed 100644 --- a/Mathlib/FieldTheory/Adjoin.lean +++ b/Mathlib/FieldTheory/Adjoin.lean @@ -452,7 +452,7 @@ theorem lift_bot (K : IntermediateField F E) : @[simp] theorem lift_top (K : IntermediateField F E) : - lift (F := K) ⊤ = K := by rw [lift, ←AlgHom.fieldRange_eq_map, fieldRange_val] + lift (F := K) ⊤ = K := by rw [lift, ← AlgHom.fieldRange_eq_map, fieldRange_val] @[simp] theorem adjoin_self (K : IntermediateField F E) : diff --git a/Mathlib/FieldTheory/Finite/Basic.lean b/Mathlib/FieldTheory/Finite/Basic.lean index c0f7c543d965a..f94c126b11920 100644 --- a/Mathlib/FieldTheory/Finite/Basic.lean +++ b/Mathlib/FieldTheory/Finite/Basic.lean @@ -5,6 +5,7 @@ Authors: Chris Hughes, Joey van Langen, Casper Putz -/ import Mathlib.FieldTheory.Separable import Mathlib.RingTheory.IntegralDomain +import Mathlib.Algebra.CharP.Reduced import Mathlib.Tactic.ApplyFun #align_import field_theory.finite.basic from "leanprover-community/mathlib"@"12a85fac627bea918960da036049d611b1a3ee43" @@ -156,11 +157,11 @@ theorem sum_subgroup_units_eq_zero [Ring K] [NoZeroDivisors K] ← Finset.mul_sum] at h_sum_map -- thus one of (a - 1) or ∑ G, x is zero have hzero : (((a : Kˣ) : K) - 1) = 0 ∨ ∑ x : ↥G, ((x : Kˣ) : K) = 0 := by - rw [←mul_eq_zero, sub_mul, ← h_sum_map, one_mul, sub_self] + rw [← mul_eq_zero, sub_mul, ← h_sum_map, one_mul, sub_self] apply Or.resolve_left hzero contrapose! ha ext - rwa [←sub_eq_zero] + rwa [← sub_eq_zero] /-- The sum of a subgroup of the units of a field is 1 if the subgroup is trivial and 1 otherwise -/ @[simp] @@ -207,7 +208,7 @@ theorem sum_subgroup_pow_eq_zero [CommRing K] [NoZeroDivisors K] rcases hzero with h | h · contrapose! ha ext - rw [←sub_eq_zero] + rw [← sub_eq_zero] simp_rw [SubmonoidClass.coe_pow, Units.val_pow_eq_pow_val, OneMemClass.coe_one, Units.val_one, h] · exact h diff --git a/Mathlib/FieldTheory/Finite/Polynomial.lean b/Mathlib/FieldTheory/Finite/Polynomial.lean index 8d195258422bd..a10cbb36711a9 100644 --- a/Mathlib/FieldTheory/Finite/Polynomial.lean +++ b/Mathlib/FieldTheory/Finite/Polynomial.lean @@ -93,7 +93,7 @@ theorem indicator_mem_restrictDegree (c : σ → K) : rw [mem_restrictDegree_iff_sup, indicator] intro n refine' le_trans (Multiset.count_le_of_le _ <| degrees_indicator _) (le_of_eq _) - simp_rw [← Multiset.coe_countAddMonoidHom, (Multiset.countAddMonoidHom n).map_sum, + simp_rw [← Multiset.coe_countAddMonoidHom, map_sum, AddMonoidHom.map_nsmul, Multiset.coe_countAddMonoidHom, nsmul_eq_mul, Nat.cast_id] trans refine' Finset.sum_eq_single n _ _ diff --git a/Mathlib/FieldTheory/Finite/Trace.lean b/Mathlib/FieldTheory/Finite/Trace.lean index 1706b02a38a44..6ee2bb18a4e2a 100644 --- a/Mathlib/FieldTheory/Finite/Trace.lean +++ b/Mathlib/FieldTheory/Finite/Trace.lean @@ -28,7 +28,7 @@ theorem trace_to_zmod_nondegenerate (F : Type*) [Field F] [Finite F] haveI : Fact (ringChar F).Prime := ⟨CharP.char_is_prime F _⟩ have htr := traceForm_nondegenerate (ZMod (ringChar F)) F a simp_rw [Algebra.traceForm_apply] at htr - by_contra' hf + by_contra! hf exact ha (htr hf) #align finite_field.trace_to_zmod_nondegenerate FiniteField.trace_to_zmod_nondegenerate diff --git a/Mathlib/FieldTheory/Minpoly/Basic.lean b/Mathlib/FieldTheory/Minpoly/Basic.lean index 01b36011e67c6..c66386a8854a3 100644 --- a/Mathlib/FieldTheory/Minpoly/Basic.lean +++ b/Mathlib/FieldTheory/Minpoly/Basic.lean @@ -244,7 +244,7 @@ variable [IsDomain A] [IsDomain B] theorem irreducible (hx : IsIntegral A x) : Irreducible (minpoly A x) := by refine' (irreducible_of_monic (monic hx) <| ne_one A x).2 fun f g hf hg he => _ rw [← hf.isUnit_iff, ← hg.isUnit_iff] - by_contra' h + by_contra! h have heval := congr_arg (Polynomial.aeval x) he rw [aeval A x, aeval_mul, mul_eq_zero] at heval cases' heval with heval heval diff --git a/Mathlib/FieldTheory/Perfect.lean b/Mathlib/FieldTheory/Perfect.lean index c44fefca52d7f..3a347ff4581d4 100644 --- a/Mathlib/FieldTheory/Perfect.lean +++ b/Mathlib/FieldTheory/Perfect.lean @@ -5,6 +5,7 @@ Authors: Oliver Nash -/ import Mathlib.FieldTheory.Separable import Mathlib.FieldTheory.SplittingField.Construction +import Mathlib.Algebra.CharP.Reduced /-! diff --git a/Mathlib/FieldTheory/RatFunc.lean b/Mathlib/FieldTheory/RatFunc.lean index 643ca0dd850d0..b4edc9a70cdf7 100644 --- a/Mathlib/FieldTheory/RatFunc.lean +++ b/Mathlib/FieldTheory/RatFunc.lean @@ -528,7 +528,7 @@ macro "smul_tac" : tactic => `(tactic| (first | rintro (⟨⟩ : RatFunc _) | intro) <;> - simp_rw [←ofFractionRing_smul] <;> + simp_rw [← ofFractionRing_smul] <;> simp only [add_comm, mul_comm, zero_smul, succ_nsmul, zsmul_eq_mul, mul_add, mul_one, mul_zero, neg_add, mul_neg, Int.ofNat_eq_coe, Int.cast_zero, Int.cast_add, Int.cast_one, diff --git a/Mathlib/FieldTheory/SplittingField/Construction.lean b/Mathlib/FieldTheory/SplittingField/Construction.lean index 7b8076ef4df0d..97b7af9a5d9c3 100644 --- a/Mathlib/FieldTheory/SplittingField/Construction.lean +++ b/Mathlib/FieldTheory/SplittingField/Construction.lean @@ -201,7 +201,7 @@ theorem adjoin_rootSet (n : ℕ) : have hfn0 : f ≠ 0 := by intro h; rw [h] at hndf; exact hndf rfl have hmf0 : map (algebraMap K (SplittingFieldAux n.succ f)) f ≠ 0 := map_ne_zero hfn0 rw [rootSet_def, aroots_def] - rw [algebraMap_succ, ←map_map, ←X_sub_C_mul_removeFactor _ hndf, Polynomial.map_mul] at hmf0 ⊢ + rw [algebraMap_succ, ← map_map, ← X_sub_C_mul_removeFactor _ hndf, Polynomial.map_mul] at hmf0 ⊢ -- This used to be `rw`, but we need `erw` after leanprover/lean4#2644 erw [roots_mul hmf0, Polynomial.map_sub, map_X, map_C, roots_X_sub_C, Multiset.toFinset_add, Finset.coe_union, Multiset.toFinset_singleton, Finset.coe_singleton, diff --git a/Mathlib/Geometry/Euclidean/Inversion/Calculus.lean b/Mathlib/Geometry/Euclidean/Inversion/Calculus.lean index c2a4d1240abff..a5e53fc5b547b 100644 --- a/Mathlib/Geometry/Euclidean/Inversion/Calculus.lean +++ b/Mathlib/Geometry/Euclidean/Inversion/Calculus.lean @@ -14,7 +14,7 @@ In this file we prove a formula for the derivative of `EuclideanGeometry.inversi ## Implementation notes -Since `fderiv` and related definiitons do not work for affine spaces, we deal with an inner product +Since `fderiv` and related definitions do not work for affine spaces, we deal with an inner product space in this file. ## Keywords @@ -87,8 +87,8 @@ theorem hasFDerivAt_inversion (hx : x ≠ c) : HasFDerivAt (inversion c R) ((R / dist x c) ^ 2 • (reflection (ℝ ∙ (x - c))ᗮ : F →L[ℝ] F)) x := by rcases add_left_surjective c x with ⟨x, rfl⟩ - have : HasFDerivAt (inversion c R) (_ : F →L[ℝ] F) (c + x) - · simp (config := { unfoldPartialApp := true }) only [inversion] + have : HasFDerivAt (inversion c R) (_ : F →L[ℝ] F) (c + x) := by + simp (config := { unfoldPartialApp := true }) only [inversion] simp_rw [dist_eq_norm, div_pow, div_eq_mul_inv] have A := (hasFDerivAt_id (𝕜 := ℝ) (c + x)).sub_const c have B := ((hasDerivAt_inv <| by simpa using hx).comp_hasFDerivAt _ A.norm_sq).const_mul @@ -97,8 +97,8 @@ theorem hasFDerivAt_inversion (hx : x ≠ c) : refine this.congr_fderiv (LinearMap.ext_on_codisjoint (Submodule.isCompl_orthogonal_of_completeSpace (K := ℝ ∙ x)).codisjoint (LinearMap.eqOn_span' ?_) fun y hy ↦ ?_) - · have : ((‖x‖ ^ 2) ^ 2)⁻¹ * (‖x‖ ^ 2) = (‖x‖ ^ 2)⁻¹ - · rw [← div_eq_inv_mul, sq (‖x‖ ^ 2), div_self_mul_self'] + · have : ((‖x‖ ^ 2) ^ 2)⁻¹ * (‖x‖ ^ 2) = (‖x‖ ^ 2)⁻¹ := by + rw [← div_eq_inv_mul, sq (‖x‖ ^ 2), div_self_mul_self'] simp [reflection_orthogonalComplement_singleton_eq_neg, real_inner_self_eq_norm_sq, two_mul, this, div_eq_mul_inv, mul_add, add_smul, mul_pow] · simp [Submodule.mem_orthogonal_singleton_iff_inner_right.1 hy, diff --git a/Mathlib/Geometry/Euclidean/MongePoint.lean b/Mathlib/Geometry/Euclidean/MongePoint.lean index cffee77c7eebc..7216a379ae640 100644 --- a/Mathlib/Geometry/Euclidean/MongePoint.lean +++ b/Mathlib/Geometry/Euclidean/MongePoint.lean @@ -713,7 +713,7 @@ theorem exists_dist_eq_circumradius_of_subset_insert_orthocenter {t : Triangle have h₁₂₃ := h₁₂₃ i repeat' cases' h₁₂₃ with h₁₂₃ h₁₂₃ · convert Triangle.dist_orthocenter_reflection_circumcenter t hj₂₃ - · rw [←h₂, dist_reflection_eq_of_mem _ + · rw [← h₂, dist_reflection_eq_of_mem _ (mem_affineSpan ℝ (Set.mem_image_of_mem _ (Set.mem_insert _ _)))] exact t.dist_circumcenter_eq_circumradius _ · rw [← h₃, diff --git a/Mathlib/Geometry/Manifold/ChartedSpace.lean b/Mathlib/Geometry/Manifold/ChartedSpace.lean index 5db698d70957e..4b39f2078907c 100644 --- a/Mathlib/Geometry/Manifold/ChartedSpace.lean +++ b/Mathlib/Geometry/Manifold/ChartedSpace.lean @@ -361,9 +361,9 @@ def Pregroupoid.groupoid (PG : Pregroupoid H) : StructureGroupoid H where trans' e e' he he' := by constructor · apply PG.comp he.1 he'.1 e.open_source e'.open_source - apply e.continuous_toFun.preimage_open_of_open e.open_source e'.open_source + apply e.continuous_toFun.isOpen_inter_preimage e.open_source e'.open_source · apply PG.comp he'.2 he.2 e'.open_target e.open_target - apply e'.continuous_invFun.preimage_open_of_open e'.open_target e.open_target + apply e'.continuous_invFun.isOpen_inter_preimage e'.open_target e.open_target symm' e he := ⟨he.2, he.1⟩ id_mem' := ⟨PG.id_mem, PG.id_mem⟩ locality' e he := by @@ -377,7 +377,7 @@ def Pregroupoid.groupoid (PG : Pregroupoid H) : StructureGroupoid H where · refine' PG.locality e.open_target fun x xu ↦ _ rcases he (e.symm x) (e.map_target xu) with ⟨s, s_open, xs, hs⟩ refine' ⟨e.target ∩ e.symm ⁻¹' s, _, ⟨xu, xs⟩, _⟩ - · exact ContinuousOn.preimage_open_of_open e.continuous_invFun e.open_target s_open + · exact ContinuousOn.isOpen_inter_preimage e.continuous_invFun e.open_target s_open · rw [← inter_assoc, inter_self] convert hs.2 using 1 dsimp [LocalHomeomorph.restr] @@ -896,7 +896,7 @@ protected def localHomeomorph (e : LocalEquiv M H) (he : e ∈ c.atlas) : exact ⟨e, he, ⟨s, s_open, rfl⟩⟩ continuous_invFun := by letI : TopologicalSpace M := c.toTopologicalSpace - apply continuousOn_open_of_generateFrom + apply continuousOn_isOpen_of_generateFrom intro t ht simp only [exists_prop, mem_iUnion, mem_singleton_iff] at ht rcases ht with ⟨e', e'_atlas, s, s_open, ts⟩ @@ -1019,7 +1019,7 @@ theorem StructureGroupoid.compatible_of_mem_maximalAtlas {e e' : LocalHomeomorph set f := chartAt (H := H) (e.symm x) let s := e.target ∩ e.symm ⁻¹' f.source have hs : IsOpen s := by - apply e.symm.continuous_toFun.preimage_open_of_open <;> apply open_source + apply e.symm.continuous_toFun.isOpen_inter_preimage <;> apply open_source have xs : x ∈ s := by simp only [mem_inter_iff, mem_preimage, mem_chart_source, and_true] exact ((mem_inter_iff _ _ _).1 hx).1 @@ -1159,7 +1159,7 @@ protected instance instHasGroupoid [ClosedUnderRestriction G] : HasGroupoid s G refine' G.eq_on_source _ (subtypeRestr_symm_trans_subtypeRestr s (chartAt H x) (chartAt H x')) apply closedUnderRestriction' · exact G.compatible (chart_mem_atlas _ _) (chart_mem_atlas _ _) - · exact preimage_open_of_open_symm (chartAt _ _) s.2 + · exact isOpen_inter_preimage_symm (chartAt _ _) s.2 #align topological_space.opens.has_groupoid TopologicalSpace.Opens.instHasGroupoid theorem chartAt_inclusion_symm_eventuallyEq {U V : Opens M} (hUV : U ≤ V) {x : U} : @@ -1236,7 +1236,7 @@ def Structomorph.trans (e : Structomorph G M M') (e' : Structomorph G M' M'') : have hg₂ := mem_chart_source (H := H) y let s := (c.symm ≫ₕ f₁).source ∩ c.symm ≫ₕ f₁ ⁻¹' g.source have open_s : IsOpen s := by - apply (c.symm ≫ₕ f₁).continuous_toFun.preimage_open_of_open <;> apply open_source + apply (c.symm ≫ₕ f₁).continuous_toFun.isOpen_inter_preimage <;> apply open_source have : x ∈ s := by constructor · simp only [trans_source, preimage_univ, inter_univ, Homeomorph.toLocalHomeomorph_source] diff --git a/Mathlib/Geometry/Manifold/Complex.lean b/Mathlib/Geometry/Manifold/Complex.lean index 03e6a15d0dba7..a34ebf30f1158 100644 --- a/Mathlib/Geometry/Manifold/Complex.lean +++ b/Mathlib/Geometry/Manifold/Complex.lean @@ -100,7 +100,7 @@ theorem norm_eqOn_of_isPreconnected_of_isMaxOn {f : M → F} {U : Set M} {c : M} (Eq.trans · hx.2) have hVne : (U ∩ V).Nonempty := ⟨c, hcU, hcU, rfl⟩ set W := U ∩ {z | ‖f z‖ = ‖f c‖}ᶜ - have hWo : IsOpen W := hd.continuousOn.norm.preimage_open_of_open ho isOpen_ne + have hWo : IsOpen W := hd.continuousOn.norm.isOpen_inter_preimage ho isOpen_ne have hdVW : Disjoint V W := disjoint_compl_right.mono inf_le_right inf_le_right have hUVW : U ⊆ V ∪ W := fun x hx => (eq_or_ne ‖f x‖ ‖f c‖).imp (.intro hx) (.intro hx) exact hc.subset_left_of_subset_union hVo hWo hdVW hUVW hVne diff --git a/Mathlib/Geometry/Manifold/Instances/Sphere.lean b/Mathlib/Geometry/Manifold/Instances/Sphere.lean index c45388ea58b2b..93e426d1b8490 100644 --- a/Mathlib/Geometry/Manifold/Instances/Sphere.lean +++ b/Mathlib/Geometry/Manifold/Instances/Sphere.lean @@ -229,7 +229,7 @@ theorem stereo_left_inv (hv : ‖v‖ = 1) {x : sphere (0 : E) 1} (hx : (x : E) · exact sq _ -- Porting note : added to work around cancel_denoms and nlinarith failures have duh : ‖y.val‖ ^ 2 = 1 - a ^ 2 := by - rw [←Submodule.coe_norm, pythag]; ring + rw [← Submodule.coe_norm, pythag]; ring -- two facts which will be helpful for clearing denominators in the main calculation have ha : 1 - a ≠ 0 := by have : a < 1 := (inner_lt_one_iff_real_of_norm_one hv (by simp)).mpr hx.symm @@ -289,7 +289,7 @@ def stereographic (hv : ‖v‖ = 1) : LocalHomeomorph (sphere (0 : E) 1) (ℝ map_source' := by simp map_target' {w} _ := fun h => (stereoInvFun_ne_north_pole hv w) (Set.eq_of_mem_singleton h) left_inv' x hx := stereo_left_inv hv fun h => hx (by - rw [←h] at hv + rw [← h] at hv apply Subtype.ext dsimp exact h) @@ -511,7 +511,7 @@ theorem range_mfderiv_coe_sphere {n : ℕ} [Fact (finrank ℝ E = n + 1)] (v : s rw [((contMDiff_coe_sphere v).mdifferentiableAt le_top).mfderiv] dsimp [chartAt] -- rw [LinearIsometryEquiv.toHomeomorph_symm] - -- rw [←LinearIsometryEquiv.coe_toHomeomorph] + -- rw [← LinearIsometryEquiv.coe_toHomeomorph] simp only [chartAt, stereographic_neg_apply, fderivWithin_univ, LinearIsometryEquiv.toHomeomorph_symm, LinearIsometryEquiv.coe_toHomeomorph, LinearIsometryEquiv.map_zero, mfld_simps] diff --git a/Mathlib/Geometry/Manifold/PartitionOfUnity.lean b/Mathlib/Geometry/Manifold/PartitionOfUnity.lean index 7d55c8b1d8b63..059b502b06fe5 100644 --- a/Mathlib/Geometry/Manifold/PartitionOfUnity.lean +++ b/Mathlib/Geometry/Manifold/PartitionOfUnity.lean @@ -159,8 +159,7 @@ theorem sum_eq_one {x} (hx : x ∈ s) : ∑ᶠ i, f i x = 1 := #align smooth_partition_of_unity.sum_eq_one SmoothPartitionOfUnity.sum_eq_one theorem exists_pos_of_mem {x} (hx : x ∈ s) : ∃ i, 0 < f i x := by - by_contra h - push_neg at h + by_contra! h have H : ∀ i, f i x = 0 := fun i ↦ le_antisymm (h i) (f.nonneg i x) have := f.sum_eq_one hx simp_rw [H] at this @@ -643,7 +642,7 @@ theorem IsOpen.exists_msmooth_support_eq (hs : IsOpen s) : Smooth I 𝓘(ℝ) g ∧ Set.range g ⊆ Set.Icc 0 1 := by intro i apply IsOpen.exists_msmooth_support_eq_aux - exact LocalHomeomorph.preimage_open_of_open_symm _ hs + exact LocalHomeomorph.isOpen_inter_preimage_symm _ hs choose g g_supp g_diff hg using A have h'g : ∀ c x, 0 ≤ g c x := fun c x ↦ (hg c (mem_range_self (f := g c) x)).1 have h''g : ∀ c x, 0 ≤ f c x * g c (chartAt H c x) := diff --git a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean index 36c349057386d..40c7bb08fe3ac 100644 --- a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean +++ b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean @@ -1052,7 +1052,7 @@ theorem extend_symm_continuousWithinAt_comp_right_iff {X} [TopologicalSpace X] { theorem isOpen_extend_preimage' {s : Set E} (hs : IsOpen s) : IsOpen ((f.extend I).source ∩ f.extend I ⁻¹' s) := - (continuousOn_extend f I).preimage_open_of_open (isOpen_extend_source _ _) hs + (continuousOn_extend f I).isOpen_inter_preimage (isOpen_extend_source _ _) hs #align local_homeomorph.is_open_extend_preimage' LocalHomeomorph.isOpen_extend_preimage' theorem isOpen_extend_preimage {s : Set E} (hs : IsOpen s) : IsOpen (f.source ∩ f.extend I ⁻¹' s) := diff --git a/Mathlib/Geometry/RingedSpace/Basic.lean b/Mathlib/Geometry/RingedSpace/Basic.lean index 4b0cfefc197d2..087b43f3bcc11 100644 --- a/Mathlib/Geometry/RingedSpace/Basic.lean +++ b/Mathlib/Geometry/RingedSpace/Basic.lean @@ -75,7 +75,7 @@ theorem isUnit_res_of_isUnit_germ (U : Opens X) (f : X.presheaf.obj (op U)) (x : obtain ⟨W', hxW', i₁, i₂, heq'⟩ := X.presheaf.germ_eq x.1 hxW hxW _ _ heq use W', i₁ ≫ Opens.infLELeft U V, hxW' rw [(X.presheaf.map i₂.op).map_one, (X.presheaf.map i₁.op).map_mul] at heq' - rw [← comp_apply, ←X.presheaf.map_comp, ←comp_apply, ←X.presheaf.map_comp, ←op_comp] at heq' + rw [← comp_apply, ← X.presheaf.map_comp, ← comp_apply, ← X.presheaf.map_comp, ← op_comp] at heq' exact isUnit_of_mul_eq_one _ _ heq' set_option linter.uppercaseLean3 false in #align algebraic_geometry.RingedSpace.is_unit_res_of_is_unit_germ AlgebraicGeometry.RingedSpace.isUnit_res_of_isUnit_germ @@ -100,7 +100,7 @@ theorem isUnit_of_isUnit_germ (U : Opens X) (f : X.presheaf.obj (op U)) erw [germ_res_apply, germ_res_apply] apply (IsUnit.mul_right_inj (h ⟨z, (iVU x).le hzVx⟩)).mp -- Porting note : now need explicitly typing the rewrites - rw [←show X.presheaf.germ ⟨z, hzVx⟩ (X.presheaf.map (iVU x).op f) = + rw [← show X.presheaf.germ ⟨z, hzVx⟩ (X.presheaf.map (iVU x).op f) = X.presheaf.germ ⟨z, ((iVU x) ⟨z, hzVx⟩).2⟩ f from X.presheaf.germ_res_apply (iVU x) ⟨z, hzVx⟩ f] -- Porting note : change was not necessary in Lean3 diff --git a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean index 40c18aea1f299..7999874f04cd4 100644 --- a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean +++ b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean @@ -148,7 +148,7 @@ instance coequalizer_π_app_isLocalRingHom have := ι_comp_coequalizerComparison f.1 g.1 SheafedSpace.forgetToPresheafedSpace rw [← PreservesCoequalizer.iso_hom] at this erw [SheafedSpace.congr_app this.symm (op U)] - rw [PresheafedSpace.comp_c_app, ←PresheafedSpace.colimitPresheafObjIsoComponentwiseLimit_hom_π] + rw [PresheafedSpace.comp_c_app, ← PresheafedSpace.colimitPresheafObjIsoComponentwiseLimit_hom_π] -- Porting note : this instance has to be manually added haveI : IsIso (PreservesCoequalizer.iso SheafedSpace.forgetToPresheafedSpace f.val g.val).hom.c := PresheafedSpace.c_isIso_of_iso _ @@ -215,7 +215,7 @@ theorem imageBasicOpen_image_preimage : theorem imageBasicOpen_image_open : IsOpen ((coequalizer.π f.1 g.1).base '' (imageBasicOpen f g U s).1) := by - rw [←(TopCat.homeoOfIso (PreservesCoequalizer.iso (SheafedSpace.forget _) f.1 + rw [← (TopCat.homeoOfIso (PreservesCoequalizer.iso (SheafedSpace.forget _) f.1 g.1)).isOpen_preimage, TopCat.coequalizer_isOpen_iff, ← Set.preimage_comp] erw [← coe_comp] rw [PreservesCoequalizer.iso_hom, ι_comp_coequalizerComparison] diff --git a/Mathlib/Geometry/RingedSpace/OpenImmersion.lean b/Mathlib/Geometry/RingedSpace/OpenImmersion.lean index 6b7d9c7651e49..dbdf4c3b2f63e 100644 --- a/Mathlib/Geometry/RingedSpace/OpenImmersion.lean +++ b/Mathlib/Geometry/RingedSpace/OpenImmersion.lean @@ -892,7 +892,7 @@ theorem image_preimage_is_empty (j : Discrete ι) (h : i ≠ j) (U : Opens (F.ob simp_rw [CategoryTheory.Iso.trans_hom, ← TopCat.comp_app, ← PresheafedSpace.comp_base] at eq rw [ι_preservesColimitsIso_inv] at eq -- Porting note : without this `erw`, change does not work - erw [←comp_apply, ←comp_apply] at eq + erw [← comp_apply, ← comp_apply] at eq change ((SheafedSpace.forget C).map (colimit.ι F i) ≫ _) y = ((SheafedSpace.forget C).map (colimit.ι F j) ≫ _) x at eq diff --git a/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean b/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean index 9bc2b4f1470c1..58d198bc3bf29 100644 --- a/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean +++ b/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean @@ -191,7 +191,7 @@ theorem snd_invApp_t_app' (i j k : D.J) (U : Opens (pullback (D.f i j) (D.f i k) dsimp only [Functor.op, Opens.map, IsOpenMap.functor, unop_op, Opens.coe_mk] congr have := (𝖣.t_fac k i j).symm - rw [←IsIso.inv_comp_eq] at this + rw [← IsIso.inv_comp_eq] at this replace this := (congr_arg ((PresheafedSpace.Hom.base ·)) this).symm replace this := congr_arg (ContinuousMap.toFun ·) this dsimp at this @@ -201,14 +201,14 @@ theorem snd_invApp_t_app' (i j k : D.J) (U : Opens (pullback (D.f i j) (D.f i k) erw [this, Set.image_comp, Set.image_comp, Set.preimage_image_eq] swap · refine Function.HasLeftInverse.injective ⟨(D.t i k).base, fun x => ?_⟩ - rw [←comp_apply, ←comp_base, D.t_inv, id_base, id_apply] + rw [← comp_apply, ← comp_base, D.t_inv, id_base, id_apply] refine congr_arg (_ '' ·) ?_ refine congr_fun ?_ _ refine Set.image_eq_preimage_of_inverse ?_ ?_ · intro x - rw [←comp_apply, ←comp_base, IsIso.inv_hom_id, id_base, id_apply] + rw [← comp_apply, ← comp_base, IsIso.inv_hom_id, id_base, id_apply] · intro x - rw [←comp_apply, ←comp_base, IsIso.hom_inv_id, id_base, id_apply] + rw [← comp_apply, ← comp_base, IsIso.hom_inv_id, id_base, id_apply] · rw [← IsIso.eq_inv_comp, IsOpenImmersion.inv_invApp, Category.assoc, (D.t' k i j).c.naturality_assoc] simp_rw [← Category.assoc] @@ -267,7 +267,7 @@ theorem ι_image_preimage_eq (i j : D.J) (U : Opens (D.U i).carrier) : change (D.t i j ≫ D.t j i).base '' _ = _ rw [𝖣.t_inv] simp - · rw [←coe_comp, ← TopCat.mono_iff_injective] + · rw [← coe_comp, ← TopCat.mono_iff_injective] infer_instance #align algebraic_geometry.PresheafedSpace.glue_data.ι_image_preimage_eq AlgebraicGeometry.PresheafedSpace.GlueData.ι_image_preimage_eq @@ -426,9 +426,9 @@ theorem ιInvApp_π {i : D.J} (U : Opens (D.U i).carrier) : · exact h2.symm · have := D.ι_gluedIso_inv (PresheafedSpace.forget _) i dsimp at this - rw [←this, coe_comp] + rw [← this, coe_comp] refine Function.Injective.comp ?_ (TopCat.GlueData.ι_injective D.toTopGlueData i) - rw [←TopCat.mono_iff_injective] + rw [← TopCat.mono_iff_injective] infer_instance delta ιInvApp rw [limit.lift_π] diff --git a/Mathlib/Geometry/RingedSpace/PresheafedSpace/HasColimits.lean b/Mathlib/Geometry/RingedSpace/PresheafedSpace/HasColimits.lean index cc6f56357a4a2..4c2508bf6ad00 100644 --- a/Mathlib/Geometry/RingedSpace/PresheafedSpace/HasColimits.lean +++ b/Mathlib/Geometry/RingedSpace/PresheafedSpace/HasColimits.lean @@ -220,7 +220,7 @@ def colimitCocone (F : J ⥤ PresheafedSpace.{_, _, v} C) : Cocone F where NatTrans.comp_app, Functor.leftOp_map, pushforwardDiagramToColimit_map] dsimp rw [NatTrans.comp_app, NatTrans.comp_app, pushforwardEq_hom_app, id.def, eqToHom_op, - Pushforward.comp_inv_app, id_comp, pushforwardMap_app, ←assoc] + Pushforward.comp_inv_app, id_comp, pushforwardMap_app, ← assoc] congr 1 } set_option linter.uppercaseLean3 false in #align algebraic_geometry.PresheafedSpace.colimit_cocone AlgebraicGeometry.PresheafedSpace.colimitCocone @@ -394,7 +394,7 @@ def colimitPresheafObjIsoComponentwiseLimit (F : J ⥤ PresheafedSpace.{_, _, v} simp only [Functor.op_obj, unop_op, op_inj_iff, Opens.map_coe, SetLike.ext'_iff, Set.preimage_preimage] refine congr_arg (Set.preimage . U.1) (funext fun x => ?_) - erw [←comp_app] + erw [← comp_app] congr exact ι_preservesColimitsIso_inv (forget C) F (unop X) · intro X Y f @@ -403,7 +403,7 @@ def colimitPresheafObjIsoComponentwiseLimit (F : J ⥤ PresheafedSpace.{_, _, v} erw [Category.id_comp] rw [Category.assoc] erw [← (F.obj (unop Y)).presheaf.map_comp, (F.map f.unop).c.naturality_assoc, - ←(F.obj (unop Y)).presheaf.map_comp] + ← (F.obj (unop Y)).presheaf.map_comp] rfl set_option linter.uppercaseLean3 false in #align algebraic_geometry.PresheafedSpace.colimit_presheaf_obj_iso_componentwise_limit AlgebraicGeometry.PresheafedSpace.colimitPresheafObjIsoComponentwiseLimit diff --git a/Mathlib/GroupTheory/ClassEquation.lean b/Mathlib/GroupTheory/ClassEquation.lean index 4b013749c1999..dec8cca31a06c 100644 --- a/Mathlib/GroupTheory/ClassEquation.lean +++ b/Mathlib/GroupTheory/ClassEquation.lean @@ -43,7 +43,7 @@ theorem Group.sum_card_conj_classes_eq_card [Finite G] : ∑ᶠ x : ConjClasses G, x.carrier.ncard = Nat.card G := by classical cases nonempty_fintype G - rw [Nat.card_eq_fintype_card, ←sum_conjClasses_card_eq_card, finsum_eq_sum_of_fintype] + rw [Nat.card_eq_fintype_card, ← sum_conjClasses_card_eq_card, finsum_eq_sum_of_fintype] simp [Set.ncard_eq_toFinset_card'] /-- The **class equation** for finite groups. The cardinality of a group is equal to the size @@ -63,7 +63,7 @@ theorem Group.nat_card_center_add_sum_card_noncenter_eq_card [Finite G] : Fintype.card (Subgroup.center G) = Fintype.card ((noncenter G)ᶜ : Set _) := Fintype.card_congr ((mk_bijOn G).equiv _) _ = Finset.card (Finset.univ \ (noncenter G).toFinset) := - by rw [←Set.toFinset_card, Set.toFinset_compl, Finset.compl_eq_univ_sdiff] + by rw [← Set.toFinset_card, Set.toFinset_compl, Finset.compl_eq_univ_sdiff] _ = _ := ?_ rw [Finset.card_eq_sum_ones] refine Finset.sum_congr rfl ?_ @@ -79,6 +79,7 @@ theorem Group.card_center_add_sum_card_noncenter_eq_card (G) [Group G] ∑ x in (noncenter G).toFinset, x.carrier.toFinset.card = Fintype.card G := by convert Group.nat_card_center_add_sum_card_noncenter_eq_card G using 2 · simp - · rw [←finsum_set_coe_eq_finsum_mem (noncenter G), finsum_eq_sum_of_fintype, ←Finset.sum_set_coe] + · rw [← finsum_set_coe_eq_finsum_mem (noncenter G), finsum_eq_sum_of_fintype, + ← Finset.sum_set_coe] simp · simp diff --git a/Mathlib/GroupTheory/CommutingProbability.lean b/Mathlib/GroupTheory/CommutingProbability.lean index 11b4a6ef32cc0..d26941786e506 100644 --- a/Mathlib/GroupTheory/CommutingProbability.lean +++ b/Mathlib/GroupTheory/CommutingProbability.lean @@ -49,16 +49,16 @@ theorem commProb_def : #align comm_prob_def commProb_def theorem commProb_prod (M' : Type*) [Mul M'] : commProb (M × M') = commProb M * commProb M' := by - simp_rw [commProb_def, div_mul_div_comm, Nat.card_prod, Nat.cast_mul, mul_pow, ←Nat.cast_mul, - ←Nat.card_prod, Commute, SemiconjBy, Prod.ext_iff] + simp_rw [commProb_def, div_mul_div_comm, Nat.card_prod, Nat.cast_mul, mul_pow, ← Nat.cast_mul, + ← Nat.card_prod, Commute, SemiconjBy, Prod.ext_iff] congr 2 exact Nat.card_congr ⟨fun x => ⟨⟨⟨x.1.1.1, x.1.2.1⟩, x.2.1⟩, ⟨⟨x.1.1.2, x.1.2.2⟩, x.2.2⟩⟩, fun x => ⟨⟨⟨x.1.1.1, x.2.1.1⟩, ⟨x.1.1.2, x.2.1.2⟩⟩, ⟨x.1.2, x.2.2⟩⟩, fun x => rfl, fun x => rfl⟩ theorem commProb_pi (i : α → Type*) [Fintype α] [∀ a, Mul (i a)] : commProb (∀ a, i a) = ∏ a, commProb (i a) := by - simp_rw [commProb_def, Finset.prod_div_distrib, Finset.prod_pow, ←Nat.cast_prod, - ←Nat.card_pi, Commute, SemiconjBy, Function.funext_iff] + simp_rw [commProb_def, Finset.prod_div_distrib, Finset.prod_pow, ← Nat.cast_prod, + ← Nat.card_pi, Commute, SemiconjBy, Function.funext_iff] congr 2 exact Nat.card_congr ⟨fun x a => ⟨⟨x.1.1 a, x.1.2 a⟩, x.2 a⟩, fun x => ⟨⟨fun a => (x a).1.1, fun a => (x a).1.2⟩, fun a => (x a).2⟩, fun x => rfl, fun x => rfl⟩ diff --git a/Mathlib/GroupTheory/Complement.lean b/Mathlib/GroupTheory/Complement.lean index 13de784ede750..08d1f9bdb4f39 100644 --- a/Mathlib/GroupTheory/Complement.lean +++ b/Mathlib/GroupTheory/Complement.lean @@ -358,7 +358,7 @@ theorem equiv_fst_eq_iff_leftCosetEquivalence {g₁ g₂ : G} : rw [LeftCosetEquivalence, leftCoset_eq_iff] constructor · intro h - rw [← hSK.equiv_fst_mul_equiv_snd g₂, ←hSK.equiv_fst_mul_equiv_snd g₁, ← h, + rw [← hSK.equiv_fst_mul_equiv_snd g₂, ← hSK.equiv_fst_mul_equiv_snd g₁, ← h, mul_inv_rev, ← mul_assoc, inv_mul_cancel_right, ← coe_inv, ← coe_mul] exact Subtype.property _ · intro h @@ -374,7 +374,7 @@ theorem equiv_snd_eq_iff_rightCosetEquivalence {g₁ g₂ : G} : rw [RightCosetEquivalence, rightCoset_eq_iff] constructor · intro h - rw [← hHT.equiv_fst_mul_equiv_snd g₂, ←hHT.equiv_fst_mul_equiv_snd g₁, ← h, + rw [← hHT.equiv_fst_mul_equiv_snd g₂, ← hHT.equiv_fst_mul_equiv_snd g₁, ← h, mul_inv_rev, mul_assoc, mul_inv_cancel_left, ← coe_inv, ← coe_mul] exact Subtype.property _ · intro h diff --git a/Mathlib/GroupTheory/Exponent.lean b/Mathlib/GroupTheory/Exponent.lean index cea2702876a3c..1479a4106e6b8 100644 --- a/Mathlib/GroupTheory/Exponent.lean +++ b/Mathlib/GroupTheory/Exponent.lean @@ -155,7 +155,7 @@ theorem exponent_min' (n : ℕ) (hpos : 0 < n) (hG : ∀ g : G, g ^ n = 1) : exp @[to_additive] theorem exponent_min (m : ℕ) (hpos : 0 < m) (hm : m < exponent G) : ∃ g : G, g ^ m ≠ 1 := by - by_contra' h + by_contra! h have hcon : exponent G ≤ m := exponent_min' m hpos h linarith #align monoid.exponent_min Monoid.exponent_min @@ -256,7 +256,7 @@ theorem exponent_ne_zero_iff_range_orderOf_finite (h : ∀ g : G, 0 < orderOf g) rw [h, zero_dvd_iff] at this exact htpos.ne' this refine' exponent_dvd_of_forall_pow_eq_one _ _ fun g => _ - rw [←pow_mod_orderOf, Nat.mod_eq_zero_of_dvd, pow_zero g] + rw [← pow_mod_orderOf, Nat.mod_eq_zero_of_dvd, pow_zero g] apply Finset.dvd_prod_of_mem rw [← Finset.mem_coe, ht] exact Set.mem_range_self g @@ -314,7 +314,7 @@ theorem exponent_eq_iSup_orderOf (h : ∀ g : G, 0 < orderOf g) : apply order_dvd_exponent refine' Nat.dvd_of_factors_subperm he _ rw [List.subperm_ext_iff] - by_contra' h + by_contra! h obtain ⟨p, hp, hpe⟩ := h replace hp := Nat.prime_of_mem_factors hp simp only [Nat.factors_count_eq] at hpe @@ -428,7 +428,7 @@ theorem MonoidHom.exponent_dvd {F M₁ M₂ : Type*} [Monoid M₁] [Monoid M₂] {f : F} (hf : Function.Surjective f) : exponent M₂ ∣ exponent M₁ := by refine Monoid.exponent_dvd_of_forall_pow_eq_one M₂ _ fun m₂ ↦ ?_ obtain ⟨m₁, rfl⟩ := hf m₂ - rw [←map_pow, pow_exponent_eq_one, map_one] + rw [← map_pow, pow_exponent_eq_one, map_one] /-- The exponent of finite product of monoids is the `Finset.lcm` of the exponents of the constituent monoids. -/ diff --git a/Mathlib/GroupTheory/GroupAction/Basic.lean b/Mathlib/GroupTheory/GroupAction/Basic.lean index 2045b9d4f69b3..119e7e716ec72 100644 --- a/Mathlib/GroupTheory/GroupAction/Basic.lean +++ b/Mathlib/GroupTheory/GroupAction/Basic.lean @@ -385,7 +385,7 @@ theorem orbitRel.Quotient.mem_orbit {a : α} {x : orbitRel.Quotient G α} : #align add_action.orbit_rel.quotient.mem_orbit AddAction.orbitRel.Quotient.mem_orbit /-- Note that `hφ = Quotient.out_eq'` is a useful choice here. -/ -@[to_additive "Note that `hφ = quotient.out_eq'` is m useful choice here."] +@[to_additive "Note that `hφ = Quotient.out_eq'` is a useful choice here."] theorem orbitRel.Quotient.orbit_eq_orbit_out (x : orbitRel.Quotient G α) {φ : orbitRel.Quotient G α → α} (hφ : letI := orbitRel G α; RightInverse φ Quotient.mk') : orbitRel.Quotient.orbit x = MulAction.orbit G (φ x) := by diff --git a/Mathlib/GroupTheory/GroupAction/BigOperators.lean b/Mathlib/GroupTheory/GroupAction/BigOperators.lean index c919d472a35eb..5d3ab0dc7c357 100644 --- a/Mathlib/GroupTheory/GroupAction/BigOperators.lean +++ b/Mathlib/GroupTheory/GroupAction/BigOperators.lean @@ -51,7 +51,7 @@ theorem Multiset.smul_sum {r : α} {s : Multiset β} : r • s.sum = (s.map ((· theorem Finset.smul_sum {r : α} {f : γ → β} {s : Finset γ} : (r • ∑ x in s, f x) = ∑ x in s, r • f x := - (DistribSMul.toAddMonoidHom β r).map_sum f s + map_sum (DistribSMul.toAddMonoidHom β r) f s #align finset.smul_sum Finset.smul_sum end @@ -66,7 +66,7 @@ theorem Multiset.smul_prod {r : α} {s : Multiset β} : r • s.prod = (s.map (( theorem Finset.smul_prod {r : α} {f : γ → β} {s : Finset γ} : (r • ∏ x in s, f x) = ∏ x in s, r • f x := - (MulDistribMulAction.toMonoidHom β r).map_prod f s + map_prod (MulDistribMulAction.toMonoidHom β r) f s #align finset.smul_prod Finset.smul_prod end diff --git a/Mathlib/GroupTheory/GroupAction/Quotient.lean b/Mathlib/GroupTheory/GroupAction/Quotient.lean index 743d222b58bb4..290746f5cf309 100644 --- a/Mathlib/GroupTheory/GroupAction/Quotient.lean +++ b/Mathlib/GroupTheory/GroupAction/Quotient.lean @@ -357,7 +357,7 @@ theorem ConjClasses.card_carrier [Group G] [Fintype G] (g : G) [Fintype (ConjCla Fintype.card G / Fintype.card (MulAction.stabilizer (ConjAct G) g) := by classical rw [Fintype.card_congr <| ConjAct.toConjAct (G := G) |>.toEquiv] - rw [←MulAction.card_orbit_mul_card_stabilizer_eq_card_group (ConjAct G) g, Nat.mul_div_cancel] + rw [← MulAction.card_orbit_mul_card_stabilizer_eq_card_group (ConjAct G) g, Nat.mul_div_cancel] simp_rw [ConjAct.orbit_eq_carrier_conjClasses] exact Fintype.card_pos_iff.mpr inferInstance diff --git a/Mathlib/GroupTheory/NoncommPiCoprod.lean b/Mathlib/GroupTheory/NoncommPiCoprod.lean index 4579346e64291..a53b91f90dc75 100644 --- a/Mathlib/GroupTheory/NoncommPiCoprod.lean +++ b/Mathlib/GroupTheory/NoncommPiCoprod.lean @@ -294,7 +294,7 @@ theorem commute_subtype_of_commute (i j : ι) (hne : i ≠ j) : rintro ⟨x, hx⟩ ⟨y, hy⟩ exact hcomm i j hne x y hx hy #align subgroup.commute_subtype_of_commute Subgroup.commute_subtype_of_commute -#align add_subgroup.commute_subtype_of_commute AddSubgroup.commute_subtype_of_commute +#align add_subgroup.commute_subtype_of_commute AddSubgroup.addCommute_subtype_of_addCommute /-- The canonical homomorphism from a family of subgroups where elements from different subgroups commute -/ diff --git a/Mathlib/GroupTheory/OrderOfElement.lean b/Mathlib/GroupTheory/OrderOfElement.lean index 78d14a3c35119..c90634018c289 100644 --- a/Mathlib/GroupTheory/OrderOfElement.lean +++ b/Mathlib/GroupTheory/OrderOfElement.lean @@ -254,7 +254,7 @@ theorem orderOf_dvd_of_pow_eq_one (h : x ^ n = 1) : orderOf x ∣ n := @[to_additive] theorem orderOf_dvd_iff_pow_eq_one {n : ℕ} : orderOf x ∣ n ↔ x ^ n = 1 := - ⟨fun h => by rw [←pow_mod_orderOf, Nat.mod_eq_zero_of_dvd h, _root_.pow_zero], + ⟨fun h => by rw [← pow_mod_orderOf, Nat.mod_eq_zero_of_dvd h, _root_.pow_zero], orderOf_dvd_of_pow_eq_one⟩ #align order_of_dvd_iff_pow_eq_one orderOf_dvd_iff_pow_eq_one #align add_order_of_dvd_iff_nsmul_eq_zero addOrderOf_dvd_iff_nsmul_eq_zero @@ -308,7 +308,7 @@ theorem exists_pow_eq_self_of_coprime (h : n.Coprime (orderOf x)) : ∃ m : ℕ, by_cases h1 : orderOf x = 1 · exact ⟨0, by rw [orderOf_eq_one_iff.mp h1, one_pow, one_pow]⟩ obtain ⟨m, h⟩ := exists_mul_emod_eq_one_of_coprime h (one_lt_iff_ne_zero_and_ne_one.mpr ⟨h0, h1⟩) - exact ⟨m, by rw [← pow_mul, ←pow_mod_orderOf, h, pow_one]⟩ + exact ⟨m, by rw [← pow_mul, ← pow_mod_orderOf, h, pow_one]⟩ #align exists_pow_eq_self_of_coprime exists_pow_eq_self_of_coprime #align exists_nsmul_eq_self_of_coprime exists_nsmul_eq_self_of_coprime @@ -557,7 +557,7 @@ lemma finite_powers : (powers a : Set G).Finite ↔ IsOfFinOrder a := by obtain ⟨m, n, hmn, ha⟩ := h.exists_lt_map_eq_of_forall_mem (f := fun n : ℕ ↦ a ^ n) (fun n ↦ by simp [mem_powers_iff]) refine isOfFinOrder_iff_pow_eq_one.2 ⟨n - m, tsub_pos_iff_lt.2 hmn, ?_⟩ - rw [←mul_left_cancel_iff (a := a ^ m), ←pow_add, add_tsub_cancel_of_le hmn.le, ha, mul_one] + rw [← mul_left_cancel_iff (a := a ^ m), ← pow_add, add_tsub_cancel_of_le hmn.le, ha, mul_one] @[to_additive (attr := simp) infinite_multiples] lemma infinite_powers : (powers a : Set G).Infinite ↔ ¬ IsOfFinOrder a := finite_powers.not @@ -583,7 +583,7 @@ lemma finEquivPowers_apply (x : G) (hx) {n : Fin (orderOf x)} : @[to_additive (attr := simp, nolint simpNF) finEquivMultiples_symm_apply] lemma finEquivPowers_symm_apply (x : G) (hx) (n : ℕ) {hn : ∃ m : ℕ, x ^ m = x ^ n} : (finEquivPowers x hx).symm ⟨x ^ n, hn⟩ = ⟨n % orderOf x, Nat.mod_lt _ hx.orderOf_pos⟩ := by - rw [Equiv.symm_apply_eq, finEquivPowers_apply, Subtype.mk_eq_mk, ←pow_mod_orderOf, Fin.val_mk] + rw [Equiv.symm_apply_eq, finEquivPowers_apply, Subtype.mk_eq_mk, ← pow_mod_orderOf, Fin.val_mk] #align fin_equiv_powers_symm_apply finEquivPowers_symm_apply #align fin_equiv_multiples_symm_apply finEquivMultiples_symm_apply @@ -1037,13 +1037,13 @@ theorem Subgroup.pow_index_mem {G : Type*} [Group G] (H : Subgroup G) [Normal H] @[to_additive] theorem pow_eq_mod_card (n : ℕ) : x ^ n = x ^ (n % Fintype.card G) := by - rw [←pow_mod_orderOf, ← Nat.mod_mod_of_dvd n orderOf_dvd_card, pow_mod_orderOf] + rw [← pow_mod_orderOf, ← Nat.mod_mod_of_dvd n orderOf_dvd_card, pow_mod_orderOf] #align pow_eq_mod_card pow_eq_mod_card #align nsmul_eq_mod_card nsmul_eq_mod_card @[to_additive] theorem zpow_eq_mod_card (n : ℤ) : x ^ n = x ^ (n % Fintype.card G : ℤ) := by - rw [←zpow_mod_orderOf, ← Int.emod_emod_of_dvd n (Int.coe_nat_dvd.2 orderOf_dvd_card), + rw [← zpow_mod_orderOf, ← Int.emod_emod_of_dvd n (Int.coe_nat_dvd.2 orderOf_dvd_card), zpow_mod_orderOf] #align zpow_eq_mod_card zpow_eq_mod_card #align zsmul_eq_mod_card zsmul_eq_mod_card diff --git a/Mathlib/GroupTheory/PGroup.lean b/Mathlib/GroupTheory/PGroup.lean index ecb159284ea60..4b61b066635fd 100644 --- a/Mathlib/GroupTheory/PGroup.lean +++ b/Mathlib/GroupTheory/PGroup.lean @@ -121,7 +121,7 @@ theorem powEquiv_apply {n : ℕ} (hn : p.Coprime n) (g : G) : hG.powEquiv hn g = @[simp] theorem powEquiv_symm_apply {n : ℕ} (hn : p.Coprime n) (g : G) : - (hG.powEquiv hn).symm g = g ^ (orderOf g).gcdB n := by rw [←Nat.card_zpowers]; rfl + (hG.powEquiv hn).symm g = g ^ (orderOf g).gcdB n := by rw [← Nat.card_zpowers]; rfl #align is_p_group.pow_equiv_symm_apply IsPGroup.powEquiv_symm_apply variable [hp : Fact p.Prime] diff --git a/Mathlib/GroupTheory/Perm/Cycle/Basic.lean b/Mathlib/GroupTheory/Perm/Cycle/Basic.lean index 053d3d20e7706..15c0d11bea434 100644 --- a/Mathlib/GroupTheory/Perm/Cycle/Basic.lean +++ b/Mathlib/GroupTheory/Perm/Cycle/Basic.lean @@ -433,7 +433,7 @@ theorem IsCycle.zpowersEquivSupport_symm_apply {σ : Perm α} (hσ : IsCycle σ) #align equiv.perm.is_cycle.zpowers_equiv_support_symm_apply Equiv.Perm.IsCycle.zpowersEquivSupport_symm_apply protected theorem IsCycle.orderOf (hf : IsCycle f) : orderOf f = f.support.card := by - rw [←Fintype.card_zpowers, ← Fintype.card_coe] + rw [← Fintype.card_zpowers, ← Fintype.card_coe] convert Fintype.card_congr (IsCycle.zpowersEquivSupport hf) #align equiv.perm.is_cycle.order_of Equiv.Perm.IsCycle.orderOf diff --git a/Mathlib/GroupTheory/Perm/Cycle/Concrete.lean b/Mathlib/GroupTheory/Perm/Cycle/Concrete.lean index 01dcdb8cf0952..7bdbd0643daf3 100644 --- a/Mathlib/GroupTheory/Perm/Cycle/Concrete.lean +++ b/Mathlib/GroupTheory/Perm/Cycle/Concrete.lean @@ -307,7 +307,7 @@ theorem next_toList_eq_apply (p : Perm α) (x y : α) (hy : y ∈ toList p x) : rw [← nthLe_toList p x k (by simpa using hk)] at hk' simp_rw [← hk'] rw [next_nthLe _ (nodup_toList _ _), nthLe_toList, nthLe_toList, ← mul_apply, ← pow_succ, - length_toList, ←pow_mod_orderOf_cycleOf_apply p (k + 1), IsCycle.orderOf] + length_toList, ← pow_mod_orderOf_cycleOf_apply p (k + 1), IsCycle.orderOf] exact isCycle_cycleOf _ (mem_support.mp hy.right) #align equiv.perm.next_to_list_eq_apply Equiv.Perm.next_toList_eq_apply diff --git a/Mathlib/GroupTheory/Perm/Support.lean b/Mathlib/GroupTheory/Perm/Support.lean index 6a5437a5ab0af..d3e7fa25ec23c 100644 --- a/Mathlib/GroupTheory/Perm/Support.lean +++ b/Mathlib/GroupTheory/Perm/Support.lean @@ -438,7 +438,7 @@ theorem support_swap {x y : α} (h : x ≠ y) : support (swap x y) = {x, y} := b theorem support_swap_iff (x y : α) : support (swap x y) = {x, y} ↔ x ≠ y := by refine' ⟨fun h => _, fun h => support_swap h⟩ - by_contra' + by_contra! rw [← this] at h simp only [swap_self, support_refl, pair_eq_singleton] at h have : x ∈ ∅ := by diff --git a/Mathlib/GroupTheory/QuotientGroup.lean b/Mathlib/GroupTheory/QuotientGroup.lean index d493335e46629..4deb45abdf1ba 100644 --- a/Mathlib/GroupTheory/QuotientGroup.lean +++ b/Mathlib/GroupTheory/QuotientGroup.lean @@ -594,7 +594,7 @@ noncomputable def quotientInfEquivProdNormalQuotient (H N : Subgroup G) [N.Norma have φ_surjective : Surjective φ := fun x => x.inductionOn' <| by rintro ⟨y, hy : y ∈ (H ⊔ N)⟩; - rw [←SetLike.mem_coe] at hy + rw [← SetLike.mem_coe] at hy rw [mul_normal H N] at hy rcases hy with ⟨h, n, hh, hn, rfl⟩ use ⟨h, hh⟩ diff --git a/Mathlib/GroupTheory/SchurZassenhaus.lean b/Mathlib/GroupTheory/SchurZassenhaus.lean index d93623e351167..c81606164c27f 100644 --- a/Mathlib/GroupTheory/SchurZassenhaus.lean +++ b/Mathlib/GroupTheory/SchurZassenhaus.lean @@ -88,10 +88,10 @@ noncomputable instance : MulAction G H.QuotientDiff where theorem smul_diff' (h : H) : diff (MonoidHom.id H) α (op (h : G) • β) = diff (MonoidHom.id H) α β * h ^ H.index := by letI := H.fintypeQuotientOfFiniteIndex - rw [diff, diff, index_eq_card, ←Finset.card_univ, ←Finset.prod_const, ←Finset.prod_mul_distrib] + rw [diff, diff, index_eq_card, ← Finset.card_univ, ← Finset.prod_const, ← Finset.prod_mul_distrib] refine' Finset.prod_congr rfl fun q _ => _ simp_rw [Subtype.ext_iff, MonoidHom.id_apply, coe_mul, mul_assoc, mul_right_inj] - rw [smul_apply_eq_smul_apply_inv_smul, smul_eq_mul_unop, unop_op, mul_left_inj, ←Subtype.ext_iff, + rw [smul_apply_eq_smul_apply_inv_smul, smul_eq_mul_unop, unop_op, mul_left_inj, ← Subtype.ext_iff, Equiv.apply_eq_iff_eq, inv_smul_eq_iff] exact self_eq_mul_right.mpr ((QuotientGroup.eq_one_iff _).mpr h.2) #align subgroup.smul_diff' Subgroup.smul_diff' @@ -183,9 +183,9 @@ private theorem step1 (K : Subgroup G) (hK : K ⊔ N = ⊤) : K = ⊤ := by exact h1.coprime_dvd_left (card_comap_dvd_of_injective N K.subtype Subtype.coe_injective) obtain ⟨H, hH⟩ := h2 K h5 h6 replace hH : Fintype.card (H.map K.subtype) = N.index := by - rw [←relindex_bot_left_eq_card, ←relindex_comap, MonoidHom.comap_bot, Subgroup.ker_subtype, - relindex_bot_left, ←IsComplement'.index_eq_card (IsComplement'.symm hH), index_comap, - subtype_range, ←relindex_sup_right, hK, relindex_top_right] + rw [← relindex_bot_left_eq_card, ← relindex_comap, MonoidHom.comap_bot, Subgroup.ker_subtype, + relindex_bot_left, ← IsComplement'.index_eq_card (IsComplement'.symm hH), index_comap, + subtype_range, ← relindex_sup_right, hK, relindex_top_right] have h7 : Fintype.card N * Fintype.card (H.map K.subtype) = Fintype.card G := by rw [hH, ← N.index_mul_card, mul_comm] have h8 : (Fintype.card N).Coprime (Fintype.card (H.map K.subtype)) := by @@ -304,7 +304,7 @@ theorem exists_right_complement'_of_coprime {N : Subgroup G} [N.Normal] haveI := (Cardinal.lt_aleph0_iff_fintype.mp (lt_of_not_ge (mt Cardinal.toNat_apply_of_aleph0_le hN3))).some apply exists_right_complement'_of_coprime_of_fintype - rwa [←Nat.card_eq_fintype_card] + rwa [← Nat.card_eq_fintype_card] #align subgroup.exists_right_complement'_of_coprime Subgroup.exists_right_complement'_of_coprime /-- **Schur-Zassenhaus** for normal subgroups: diff --git a/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean b/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean index 9339d41c649fa..eaa1466940348 100644 --- a/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean +++ b/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean @@ -125,7 +125,7 @@ theorem isCyclic_of_orderOf_eq_card [Fintype α] (x : α) (hx : orderOf x = Fint classical use x simp_rw [← SetLike.mem_coe, ← Set.eq_univ_iff_forall] - rw [← Fintype.card_congr (Equiv.Set.univ α), ←Fintype.card_zpowers] at hx + rw [← Fintype.card_congr (Equiv.Set.univ α), ← Fintype.card_zpowers] at hx exact Set.eq_of_subset_of_card_le (Set.subset_univ _) (ge_of_eq hx) #align is_cyclic_of_order_of_eq_card isCyclic_of_orderOf_eq_card #align is_add_cyclic_of_order_of_eq_card isAddCyclic_of_orderOf_eq_card @@ -175,7 +175,7 @@ theorem isCyclic_of_surjective {H G F : Type*} [Group H] [Group G] [hH : IsCycli theorem orderOf_eq_card_of_forall_mem_zpowers [Fintype α] {g : α} (hx : ∀ x, x ∈ zpowers g) : orderOf g = Fintype.card α := by classical - rw [←Fintype.card_zpowers] + rw [← Fintype.card_zpowers] apply Fintype.card_of_finset' simpa using hx #align order_of_eq_card_of_forall_mem_zpowers orderOf_eq_card_of_forall_mem_zpowers @@ -188,7 +188,7 @@ theorem exists_pow_ne_one_of_isCyclic {G : Type*} [Group G] [Fintype G] [G_cycli use a contrapose! k_lt_card_G convert orderOf_le_of_pow_eq_one k_pos.bot_lt k_lt_card_G - rw [←Fintype.card_zpowers, eq_comm, Subgroup.card_eq_iff_eq_top, eq_top_iff] + rw [← Fintype.card_zpowers, eq_comm, Subgroup.card_eq_iff_eq_top, eq_top_iff] exact fun x _ ↦ ha x @[to_additive Infinite.addOrderOf_eq_zero_of_forall_mem_zmultiples] @@ -202,12 +202,12 @@ theorem Infinite.orderOf_eq_zero_of_forall_mem_zpowers [Infinite α] {g : α} Infinite.exists_not_mem_finset (Finset.image (fun x => g ^ x) <| Finset.range <| orderOf g) apply hx - rw [←ho.mem_powers_iff_mem_range_orderOf, Submonoid.mem_powers_iff] + rw [← ho.mem_powers_iff_mem_range_orderOf, Submonoid.mem_powers_iff] obtain ⟨k, hk⟩ := h x dsimp at hk obtain ⟨k, rfl | rfl⟩ := k.eq_nat_or_neg · exact ⟨k, mod_cast hk⟩ - rw [←zpow_mod_orderOf] at hk + rw [← zpow_mod_orderOf] at hk have : 0 ≤ (-k % orderOf g : ℤ) := Int.emod_nonneg (-k) (mod_cast ho.orderOf_pos.ne') refine' ⟨(-k % orderOf g : ℤ).toNat, _⟩ rwa [← zpow_ofNat, Int.toNat_of_nonneg this] @@ -233,7 +233,7 @@ instance Subgroup.isCyclic {α : Type u} [Group α] [IsCyclic α] (H : Subgroup Nat.pos_of_ne_zero fun h => hx₂ <| by rw [← hk, Int.natAbs_eq_zero.mp h, zpow_zero], by cases' k with k k - · rw [Int.ofNat_eq_coe, Int.natAbs_cast k, ← zpow_ofNat, ←Int.ofNat_eq_coe, hk] + · rw [Int.ofNat_eq_coe, Int.natAbs_cast k, ← zpow_ofNat, ← Int.ofNat_eq_coe, hk] exact hx₁ · rw [Int.natAbs_negSucc, ← Subgroup.inv_mem_iff H]; simp_all⟩ ⟨⟨⟨g ^ Nat.find hex, (Nat.find_spec hex).2⟩, fun ⟨x, hx⟩ => diff --git a/Mathlib/GroupTheory/SpecificGroups/Dihedral.lean b/Mathlib/GroupTheory/SpecificGroups/Dihedral.lean index df50ccc519fcc..89f7d4bb81a27 100644 --- a/Mathlib/GroupTheory/SpecificGroups/Dihedral.lean +++ b/Mathlib/GroupTheory/SpecificGroups/Dihedral.lean @@ -234,8 +234,8 @@ def OddCommuteEquiv (hn : Odd n) : { p : DihedralGroup n × DihedralGroup n // C simpa [sub_eq_add_neg, eq_neg_iff_add_eq_zero, hu, eq_comm (a := j) (b := 0)] using h.eq | ⟨⟨sr i, sr j⟩, h⟩ => by replace h := r.inj h - rw [←neg_sub, neg_eq_iff_add_eq_zero, hu, sub_eq_zero] at h - rw [Subtype.ext_iff, Prod.ext_iff, sr.injEq, sr.injEq, h, and_self, ←two_mul] + rw [← neg_sub, neg_eq_iff_add_eq_zero, hu, sub_eq_zero] at h + rw [Subtype.ext_iff, Prod.ext_iff, sr.injEq, sr.injEq, h, and_self, ← two_mul] exact u.inv_mul_cancel_left j right_inv := fun | .inl i => rfl @@ -253,7 +253,7 @@ lemma card_commute_odd (hn : Odd n) : lemma card_conjClasses_odd (hn : Odd n) : Nat.card (ConjClasses (DihedralGroup n)) = (n + 3) / 2 := by - rw [←Nat.mul_div_mul_left _ 2 hn.pos, ← card_commute_odd hn, mul_comm, + rw [← Nat.mul_div_mul_left _ 2 hn.pos, ← card_commute_odd hn, mul_comm, card_comm_eq_card_conjClasses_mul_card, nat_card, Nat.mul_div_left _ (mul_pos two_pos hn.pos)] diff --git a/Mathlib/GroupTheory/SpecificGroups/KleinFour.lean b/Mathlib/GroupTheory/SpecificGroups/KleinFour.lean index 6e3e5da9117b3..77713ca988654 100644 --- a/Mathlib/GroupTheory/SpecificGroups/KleinFour.lean +++ b/Mathlib/GroupTheory/SpecificGroups/KleinFour.lean @@ -192,7 +192,7 @@ def mulEquiv' (e : G₁ ≃ G₂) (he : e 1 = 1) (h : Monoid.exponent G₂ = 2) have univ₂ : {e (x * y), e x, e y, (1 : G₂)} = Finset.univ := by simpa [map_univ_equiv e, map_insert, he] using congr(Finset.map e.toEmbedding $(eq_finset_univ hx hy hxy)) - rw [←Ne.def, ←e.injective.ne_iff] at hx hy hxy + rw [← Ne.def, ← e.injective.ne_iff] at hx hy hxy rw [he] at hx hy symm apply eq_of_not_mem_of_mem_insert <| univ₂.symm ▸ mem_univ _ diff --git a/Mathlib/GroupTheory/Subgroup/Basic.lean b/Mathlib/GroupTheory/Subgroup/Basic.lean index 6337b605b9a24..a9a20a1418075 100644 --- a/Mathlib/GroupTheory/Subgroup/Basic.lean +++ b/Mathlib/GroupTheory/Subgroup/Basic.lean @@ -941,7 +941,7 @@ theorem nontrivial_iff_exists_ne_one (H : Subgroup G) : Nontrivial H ↔ ∃ x @[to_additive] theorem exists_ne_one_of_nontrivial (H : Subgroup G) [Nontrivial H] : ∃ x ∈ H, x ≠ 1 := by - rwa [←Subgroup.nontrivial_iff_exists_ne_one] + rwa [← Subgroup.nontrivial_iff_exists_ne_one] @[to_additive] theorem nontrivial_iff_ne_bot (H : Subgroup G) : Nontrivial H ↔ H ≠ ⊥ := by @@ -966,7 +966,7 @@ theorem bot_or_exists_ne_one (H : Subgroup G) : H = ⊥ ∨ ∃ x ∈ H, x ≠ ( @[to_additive] lemma ne_bot_iff_exists_ne_one {H : Subgroup G} : H ≠ ⊥ ↔ ∃ a : ↥H, a ≠ 1 := by - rw [←nontrivial_iff_ne_bot, nontrivial_iff_exists_ne_one] + rw [← nontrivial_iff_ne_bot, nontrivial_iff_exists_ne_one] simp only [ne_eq, Subtype.exists, mk_eq_one_iff, exists_prop] /-- The inf of two subgroups is their intersection. -/ @@ -3717,7 +3717,7 @@ theorem commute_of_normal_of_disjoint (H₁ H₂ : Subgroup G) (hH₁ : H₁.Nor apply H₂.mul_mem _ (H₂.inv_mem hy) apply hH₂.conj_mem _ hy #align subgroup.commute_of_normal_of_disjoint Subgroup.commute_of_normal_of_disjoint -#align add_subgroup.commute_of_normal_of_disjoint AddSubgroup.commute_of_normal_of_disjoint +#align add_subgroup.commute_of_normal_of_disjoint AddSubgroup.addCommute_of_normal_of_disjoint end SubgroupNormal diff --git a/Mathlib/GroupTheory/Submonoid/Membership.lean b/Mathlib/GroupTheory/Submonoid/Membership.lean index ca773167db2b1..d6bc8e0d2b04f 100644 --- a/Mathlib/GroupTheory/Submonoid/Membership.lean +++ b/Mathlib/GroupTheory/Submonoid/Membership.lean @@ -60,7 +60,7 @@ theorem coe_multiset_prod {M} [CommMonoid M] [SetLike B M] [SubmonoidClass B M] @[to_additive (attr := norm_cast)] --Porting note: removed `simp`, `simp` can prove it theorem coe_finset_prod {ι M} [CommMonoid M] [SetLike B M] [SubmonoidClass B M] (f : ι → S) (s : Finset ι) : ↑(∏ i in s, f i) = (∏ i in s, f i : M) := - (SubmonoidClass.subtype S : _ →* M).map_prod f s + map_prod (SubmonoidClass.subtype S) f s #align submonoid_class.coe_finset_prod SubmonoidClass.coe_finset_prod #align add_submonoid_class.coe_finset_sum AddSubmonoidClass.coe_finset_sum @@ -122,7 +122,7 @@ theorem coe_multiset_prod {M} [CommMonoid M] (S : Submonoid M) (m : Multiset S) @[to_additive (attr := norm_cast, simp)] theorem coe_finset_prod {ι M} [CommMonoid M] (S : Submonoid M) (f : ι → S) (s : Finset ι) : ↑(∏ i in s, f i) = (∏ i in s, f i : M) := - S.subtype.map_prod f s + map_prod S.subtype f s #align submonoid.coe_finset_prod Submonoid.coe_finset_prod #align add_submonoid.coe_finset_sum AddSubmonoid.coe_finset_sum @@ -232,14 +232,14 @@ theorem coe_sSup_of_directedOn {S : Set (Submonoid M)} (Sne : S.Nonempty) @[to_additive] theorem mem_sup_left {S T : Submonoid M} : ∀ {x : M}, x ∈ S → x ∈ S ⊔ T := by - rw [←SetLike.le_def] + rw [← SetLike.le_def] exact le_sup_left #align submonoid.mem_sup_left Submonoid.mem_sup_left #align add_submonoid.mem_sup_left AddSubmonoid.mem_sup_left @[to_additive] theorem mem_sup_right {S T : Submonoid M} : ∀ {x : M}, x ∈ T → x ∈ S ⊔ T := by - rw [←SetLike.le_def] + rw [← SetLike.le_def] exact le_sup_right #align submonoid.mem_sup_right Submonoid.mem_sup_right #align add_submonoid.mem_sup_right AddSubmonoid.mem_sup_right @@ -253,7 +253,7 @@ theorem mul_mem_sup {S T : Submonoid M} {x y : M} (hx : x ∈ S) (hy : y ∈ T) @[to_additive] theorem mem_iSup_of_mem {ι : Sort*} {S : ι → Submonoid M} (i : ι) : ∀ {x : M}, x ∈ S i → x ∈ iSup S := by - rw [←SetLike.le_def] + rw [← SetLike.le_def] exact le_iSup _ _ #align submonoid.mem_supr_of_mem Submonoid.mem_iSup_of_mem #align add_submonoid.mem_supr_of_mem AddSubmonoid.mem_iSup_of_mem @@ -261,7 +261,7 @@ theorem mem_iSup_of_mem {ι : Sort*} {S : ι → Submonoid M} (i : ι) : @[to_additive] theorem mem_sSup_of_mem {S : Set (Submonoid M)} {s : Submonoid M} (hs : s ∈ S) : ∀ {x : M}, x ∈ s → x ∈ sSup S := by - rw [←SetLike.le_def] + rw [← SetLike.le_def] exact le_sSup hs #align submonoid.mem_Sup_of_mem Submonoid.mem_sSup_of_mem #align add_submonoid.mem_Sup_of_mem AddSubmonoid.mem_sSup_of_mem diff --git a/Mathlib/GroupTheory/Subsemigroup/Center.lean b/Mathlib/GroupTheory/Subsemigroup/Center.lean index fc6f48af314bc..976df673386c3 100644 --- a/Mathlib/GroupTheory/Subsemigroup/Center.lean +++ b/Mathlib/GroupTheory/Subsemigroup/Center.lean @@ -185,9 +185,10 @@ theorem natCast_mem_center [NonAssocSemiring M] (n : ℕ) : (n : M) ∈ Set.cent | zero => rw [Nat.zero_eq, Nat.cast_zero, mul_zero, mul_zero, mul_zero] | succ n ihn => rw [Nat.cast_succ, mul_add, ihn, mul_add, mul_add, mul_one, mul_one] +-- See note [no_index around OfNat.ofNat] @[simp] theorem ofNat_mem_center [NonAssocSemiring M] (n : ℕ) [n.AtLeastTwo] : - OfNat.ofNat n ∈ Set.center M := + (no_index (OfNat.ofNat n)) ∈ Set.center M := natCast_mem_center M n @[simp] @@ -228,9 +229,9 @@ theorem inv_mem_center [Group M] {a : M} (ha : a ∈ Set.center M) : a⁻¹ ∈ theorem add_mem_center [Distrib M] {a b : M} (ha : a ∈ Set.center M) (hb : b ∈ Set.center M) : a + b ∈ Set.center M where comm _ := by rw [add_mul, mul_add, ha.comm, hb.comm] - left_assoc _ _ := by rw [add_mul, ha.left_assoc, hb.left_assoc, ←add_mul, ←add_mul] - mid_assoc _ _ := by rw [mul_add, add_mul, ha.mid_assoc, hb.mid_assoc, ←mul_add, ←add_mul] - right_assoc _ _ := by rw [mul_add, ha.right_assoc, hb.right_assoc, ←mul_add, ←mul_add] + left_assoc _ _ := by rw [add_mul, ha.left_assoc, hb.left_assoc, ← add_mul, ← add_mul] + mid_assoc _ _ := by rw [mul_add, add_mul, ha.mid_assoc, hb.mid_assoc, ← mul_add, ← add_mul] + right_assoc _ _ := by rw [mul_add, ha.right_assoc, hb.right_assoc, ← mul_add, ← mul_add] #align set.add_mem_center Set.add_mem_center @[simp] @@ -247,7 +248,7 @@ theorem subset_center_units [Monoid M] : ((↑) : Mˣ → M) ⁻¹' center M ⊆ fun _ ha => by rw [_root_.Semigroup.mem_center_iff] intro _ - rw [←Units.eq_iff, Units.val_mul, Units.val_mul, ha.comm] + rw [← Units.eq_iff, Units.val_mul, Units.val_mul, ha.comm] #align set.subset_center_units Set.subset_center_units #align set.subset_add_center_add_units Set.subset_addCenter_add_units diff --git a/Mathlib/GroupTheory/Sylow.lean b/Mathlib/GroupTheory/Sylow.lean index c920e1ea87d0d..830b5c9577724 100644 --- a/Mathlib/GroupTheory/Sylow.lean +++ b/Mathlib/GroupTheory/Sylow.lean @@ -612,7 +612,7 @@ theorem exists_subgroup_card_pow_succ [Fintype G] {p : ℕ} {n : ℕ} [hp : Fact rw [Set.card_image_of_injective (Subgroup.comap (mk' (H.subgroupOf H.normalizer)) (zpowers x) : Set H.normalizer) Subtype.val_injective, - pow_succ', ← hH, Fintype.card_congr hequiv, ← hx, ←Fintype.card_zpowers, ← + pow_succ', ← hH, Fintype.card_congr hequiv, ← hx, ← Fintype.card_zpowers, ← Fintype.card_prod] exact @Fintype.card_congr _ _ (_) (_) (preimageMkEquivSubgroupProdSet (H.subgroupOf H.normalizer) (zpowers x)), by diff --git a/Mathlib/GroupTheory/Transfer.lean b/Mathlib/GroupTheory/Transfer.lean index eff4dacb97672..c02ae6cc4f5e7 100644 --- a/Mathlib/GroupTheory/Transfer.lean +++ b/Mathlib/GroupTheory/Transfer.lean @@ -187,7 +187,7 @@ theorem transfer_eq_pow [FiniteIndex H] (g : G) theorem transfer_center_eq_pow [FiniteIndex (center G)] (g : G) : transfer (MonoidHom.id (center G)) g = ⟨g ^ (center G).index, (center G).pow_index_mem g⟩ := - transfer_eq_pow (id (center G)) g fun k _ hk => by rw [← mul_right_inj, ←hk.comm, + transfer_eq_pow (id (center G)) g fun k _ hk => by rw [← mul_right_inj, ← hk.comm, mul_inv_cancel_right] #align monoid_hom.transfer_center_eq_pow MonoidHom.transfer_center_eq_pow diff --git a/Mathlib/Init/Data/Nat/Bitwise.lean b/Mathlib/Init/Data/Nat/Bitwise.lean index 5d8d5c74da67e..7f67497d95a84 100644 --- a/Mathlib/Init/Data/Nat/Bitwise.lean +++ b/Mathlib/Init/Data/Nat/Bitwise.lean @@ -190,7 +190,7 @@ theorem shiftLeft'_false : ∀ n, shiftLeft' false m n = m <<< n | n + 1 => by have : 2 * (m * 2^n) = 2^(n+1)*m := by rw [Nat.mul_comm, Nat.mul_assoc, ← pow_succ]; simp - simp [shiftLeft', bit_val, shiftLeft'_false, this] + simp [shiftLeft_eq, shiftLeft', bit_val, shiftLeft'_false, this] /-- Std4 takes the unprimed name for `Nat.shiftLeft_eq m n : m <<< n = m * 2 ^ n`. -/ @[simp] @@ -199,14 +199,6 @@ lemma shiftLeft_eq' (m n : Nat) : shiftLeft m n = m <<< n := rfl @[simp] lemma shiftRight_eq (m n : Nat) : shiftRight m n = m >>> n := rfl -theorem shiftLeft_zero (m) : m <<< 0 = m := rfl - -theorem shiftLeft_succ (m n) : m <<< (n + 1) = 2 * (m <<< n) := by - simp only [shiftLeft_eq, Nat.pow_add, Nat.pow_one, ← Nat.mul_assoc, Nat.mul_comm] - -/-- `testBit m n` returns whether the `(n+1)ˢᵗ` least significant bit is `1` or `0`-/ -def testBit (m n : ℕ) : Bool := - bodd (m >>> n) #align nat.test_bit Nat.testBit lemma binaryRec_decreasing (h : n ≠ 0) : div2 n < n := by @@ -229,7 +221,7 @@ def binaryRec {C : Nat → Sort u} (z : C 0) (f : ∀ b n, C n → C (bit b n)) let n' := div2 n have _x : bit (bodd n) n' = n := by apply bit_decomp n - rw [←_x] + rw [← _x] exact f (bodd n) n' (binaryRec z f n') decreasing_by exact binaryRec_decreasing n0 #align nat.binary_rec Nat.binaryRec @@ -301,15 +293,28 @@ theorem shiftLeft_sub : ∀ (m : Nat) {n k}, k ≤ n → m <<< (n - k) = (m <<< fun _ _ _ hk => by simp only [← shiftLeft'_false, shiftLeft'_sub false _ hk] @[simp] -theorem testBit_zero (b n) : testBit (bit b n) 0 = b := - bodd_bit _ _ +theorem testBit_zero (b n) : testBit (bit b n) 0 = b := by + rw [testBit, bit] + cases b + · simp [bit0, ← Nat.mul_two] + · simp only [cond_true, bit1, bit0, shiftRight_zero, and_one_is_mod, bne_iff_ne] + simp only [← Nat.mul_two] + rw [Nat.add_mod] + simp + #align nat.test_bit_zero Nat.testBit_zero +theorem bodd_eq_and_one_ne_zero : ∀ n, bodd n = (n &&& 1 != 0) + | 0 => rfl + | 1 => rfl + | n + 2 => by simpa using bodd_eq_and_one_ne_zero n + theorem testBit_succ (m b n) : testBit (bit b n) (succ m) = testBit n m := by have : bodd (((bit b n) >>> 1) >>> m) = bodd (n >>> m) := by - dsimp [shiftRight] + simp only [shiftRight_eq_div_pow] simp [← div2_val, div2_bit] rw [← shiftRight_add, Nat.add_comm] at this + simp only [bodd_eq_and_one_ne_zero] at this exact this #align nat.test_bit_succ Nat.testBit_succ diff --git a/Mathlib/Lean/Expr/Basic.lean b/Mathlib/Lean/Expr/Basic.lean index 35237f7db112b..20350fd2aac08 100644 --- a/Mathlib/Lean/Expr/Basic.lean +++ b/Mathlib/Lean/Expr/Basic.lean @@ -445,7 +445,7 @@ def mkDirectProjection (e : Expr) (fieldName : Name) : MetaM Expr := do /-- If `e` has a structure as type with field `fieldName` (either directly or in a parent structure), `mkProjection e fieldName` creates the projection expression `e.fieldName` -/ def mkProjection (e : Expr) (fieldName : Name) : MetaM Expr := do - let .const structName _ := (← whnf (←inferType e)).getAppFn | + let .const structName _ := (← whnf (← inferType e)).getAppFn | throwError "{e} doesn't have a structure as type" let some baseStruct := findField? (← getEnv) structName fieldName | throwError "No parent of {structName} has field {fieldName}" diff --git a/Mathlib/Lean/Meta/DiscrTree.lean b/Mathlib/Lean/Meta/DiscrTree.lean index 90913a9df7895..41d79afb61159 100644 --- a/Mathlib/Lean/Meta/DiscrTree.lean +++ b/Mathlib/Lean/Meta/DiscrTree.lean @@ -43,4 +43,11 @@ partial def getSubexpressionMatches (d : DiscrTree α) (e : Expr) (config : Whnf e.foldlM (fun a f => do pure <| a ++ (← d.getSubexpressionMatches f config)) (← d.getMatch e config).reverse +/-- +Check if a `keys : Array DiscTree.Key` is "specific", +i.e. something other than `[*]` or `[=, *, *, *]`. +-/ +def keysSpecific (keys : Array DiscrTree.Key) : Bool := + keys != #[Key.star] && keys != #[Key.const `Eq 3, Key.star, Key.star, Key.star] + end Lean.Meta.DiscrTree diff --git a/Mathlib/LinearAlgebra/CliffordAlgebra/BaseChange.lean b/Mathlib/LinearAlgebra/CliffordAlgebra/BaseChange.lean index 32a5373c9f731..c7a4dce5ee57f 100644 --- a/Mathlib/LinearAlgebra/CliffordAlgebra/BaseChange.lean +++ b/Mathlib/LinearAlgebra/CliffordAlgebra/BaseChange.lean @@ -49,8 +49,8 @@ noncomputable def ofBaseChangeAux (Q : QuadraticForm R V) : CliffordAlgebra.lift Q <| by refine ⟨(ι (Q.baseChange A)).restrictScalars R ∘ₗ TensorProduct.mk R A V 1, fun v => ?_⟩ refine (CliffordAlgebra.ι_sq_scalar (Q.baseChange A) (1 ⊗ₜ v)).trans ?_ - rw [QuadraticForm.baseChange_tmul, one_mul, ←Algebra.algebraMap_eq_smul_one, - ←IsScalarTower.algebraMap_apply] + rw [QuadraticForm.baseChange_tmul, one_mul, ← Algebra.algebraMap_eq_smul_one, + ← IsScalarTower.algebraMap_apply] @[simp] theorem ofBaseChangeAux_ι (Q : QuadraticForm R V) (v : V) : ofBaseChangeAux A Q (ι Q v) = ι (Q.baseChange A) (1 ⊗ₜ v) := @@ -67,7 +67,7 @@ noncomputable def ofBaseChange (Q : QuadraticForm R V) : @[simp] theorem ofBaseChange_tmul_ι (Q : QuadraticForm R V) (z : A) (v : V) : ofBaseChange A Q (z ⊗ₜ ι Q v) = ι (Q.baseChange A) (z ⊗ₜ v) := by show algebraMap _ _ z * ofBaseChangeAux A Q (ι Q v) = ι (Q.baseChange A) (z ⊗ₜ[R] v) - rw [ofBaseChangeAux_ι, ←Algebra.smul_def, ←map_smul, TensorProduct.smul_tmul', smul_eq_mul, + rw [ofBaseChangeAux_ι, ← Algebra.smul_def, ← map_smul, TensorProduct.smul_tmul', smul_eq_mul, mul_one] @[simp] theorem ofBaseChange_tmul_one (Q : QuadraticForm R V) (z : A) : @@ -144,7 +144,7 @@ theorem toBaseChange_reverse (Q : QuadraticForm R V) (x : CliffordAlgebra (Q.bas have := FunLike.congr_fun (toBaseChange_comp_reverseOp A Q) x refine (congr_arg unop this).trans ?_; clear this refine (LinearMap.congr_fun (TensorProduct.AlgebraTensorModule.map_comp _ _ _ _).symm _).trans ?_ - rw [reverse, ←AlgEquiv.toLinearMap, ←AlgEquiv.toLinearEquiv_toLinearMap, + rw [reverse, ← AlgEquiv.toLinearMap, ← AlgEquiv.toLinearEquiv_toLinearMap, AlgEquiv.toLinearEquiv_toOpposite] dsimp -- `simp` fails here due to a timeout looking for a `Subsingleton` instance!? diff --git a/Mathlib/LinearAlgebra/CliffordAlgebra/EvenEquiv.lean b/Mathlib/LinearAlgebra/CliffordAlgebra/EvenEquiv.lean index 08c862fb3cf6f..e04848f971f0c 100644 --- a/Mathlib/LinearAlgebra/CliffordAlgebra/EvenEquiv.lean +++ b/Mathlib/LinearAlgebra/CliffordAlgebra/EvenEquiv.lean @@ -241,6 +241,7 @@ def equivEven : CliffordAlgebra Q ≃ₐ[R] CliffordAlgebra.even (Q' Q) := -- Note: times out on linting CI attribute [nolint simpNF] equivEven_symm_apply +set_option synthInstance.maxHeartbeats 30000 in /-- The representation of the clifford conjugate (i.e. the reverse of the involute) in the even subalgebra is just the reverse of the representation. -/ theorem coe_toEven_reverse_involute (x : CliffordAlgebra Q) : @@ -259,6 +260,7 @@ theorem coe_toEven_reverse_involute (x : CliffordAlgebra Q) : /-! ### Constructions needed for `CliffordAlgebra.evenEquivEvenNeg` -/ +set_option synthInstance.maxHeartbeats 30000 in /-- One direction of `CliffordAlgebra.evenEquivEvenNeg` -/ def evenToNeg (Q' : QuadraticForm R M) (h : Q' = -Q) : CliffordAlgebra.even Q →ₐ[R] CliffordAlgebra.even Q' := diff --git a/Mathlib/LinearAlgebra/Dual.lean b/Mathlib/LinearAlgebra/Dual.lean index 8f2210786a9f4..c513037785437 100644 --- a/Mathlib/LinearAlgebra/Dual.lean +++ b/Mathlib/LinearAlgebra/Dual.lean @@ -1569,7 +1569,7 @@ theorem finrank_range_dualMap_eq_finrank_range (f : V₁ →ₗ[K] V₂) : let equiv := (Subspace.quotEquivAnnihilator <| LinearMap.range f) have eq := LinearEquiv.finrank_eq (R := K) (M := (V₂ ⧸ range f)) (M₂ := { x // x ∈ Submodule.dualAnnihilator (range f) }) equiv - rw [eq, ←ker_dualMap_eq_dualAnnihilator_range] at that + rw [eq, ← ker_dualMap_eq_dualAnnihilator_range] at that -- Porting note: cannot convert at `this`? conv_rhs at that => rw [← Subspace.dual_finrank_eq] refine' add_left_injective (finrank K <| LinearMap.ker f.dualMap) _ diff --git a/Mathlib/LinearAlgebra/Finsupp.lean b/Mathlib/LinearAlgebra/Finsupp.lean index bbfaf7d4aed76..2bbafda85532f 100644 --- a/Mathlib/LinearAlgebra/Finsupp.lean +++ b/Mathlib/LinearAlgebra/Finsupp.lean @@ -826,7 +826,7 @@ variable {α} {M} {v} theorem totalOn_range (s : Set α) : LinearMap.range (Finsupp.totalOn α M R v s) = ⊤ := by rw [Finsupp.totalOn, LinearMap.range_eq_map, LinearMap.map_codRestrict, - ←LinearMap.range_le_iff_comap, range_subtype, Submodule.map_top, LinearMap.range_comp, + ← LinearMap.range_le_iff_comap, range_subtype, Submodule.map_top, LinearMap.range_comp, range_subtype] exact (span_image_eq_map_total _ _).le #align finsupp.total_on_range Finsupp.totalOn_range diff --git a/Mathlib/LinearAlgebra/FreeAlgebra.lean b/Mathlib/LinearAlgebra/FreeAlgebra.lean index d52b0834e8139..70e3130a1365c 100644 --- a/Mathlib/LinearAlgebra/FreeAlgebra.lean +++ b/Mathlib/LinearAlgebra/FreeAlgebra.lean @@ -43,7 +43,7 @@ end theorem rank_eq [CommRing R] [Nontrivial R] : Module.rank R (FreeAlgebra R X) = Cardinal.lift.{u} (Cardinal.mk (List X)) := by - rw [←(Basis.mk_eq_rank'.{_,_,_,u} (basisFreeMonoid R X)).trans (Cardinal.lift_id _), + rw [← (Basis.mk_eq_rank'.{_,_,_,u} (basisFreeMonoid R X)).trans (Cardinal.lift_id _), Cardinal.lift_umax'.{v,u}, FreeMonoid] #align free_algebra.rank_eq FreeAlgebra.rank_eq diff --git a/Mathlib/LinearAlgebra/Matrix/Charpoly/Coeff.lean b/Mathlib/LinearAlgebra/Matrix/Charpoly/Coeff.lean index d4e116a3051e0..e84e8db3e2d5d 100644 --- a/Mathlib/LinearAlgebra/Matrix/Charpoly/Coeff.lean +++ b/Mathlib/LinearAlgebra/Matrix/Charpoly/Coeff.lean @@ -250,7 +250,7 @@ theorem coeff_charpoly_mem_ideal_pow {I : Ideal R} (h : ∀ i j, M i j ∈ I) (k apply coeff_prod_mem_ideal_pow_tsub rintro i - (_ | k) · rw [Nat.zero_eq] -- porting note: `rw [Nat.zero_eq]` was not present - rw [tsub_zero, pow_one, charmatrix_apply, coeff_sub, ←smul_one_eq_diagonal, smul_apply, + rw [tsub_zero, pow_one, charmatrix_apply, coeff_sub, ← smul_one_eq_diagonal, smul_apply, smul_eq_mul, coeff_X_mul_zero, coeff_C_zero, zero_sub] apply neg_mem -- porting note: was `rw [neg_mem_iff]`, but Lean could not synth `NegMemClass` exact h (c i) i diff --git a/Mathlib/LinearAlgebra/Matrix/PosDef.lean b/Mathlib/LinearAlgebra/Matrix/PosDef.lean index 356297c13900e..0f6fc99935760 100644 --- a/Mathlib/LinearAlgebra/Matrix/PosDef.lean +++ b/Mathlib/LinearAlgebra/Matrix/PosDef.lean @@ -197,7 +197,7 @@ noncomputable def NormedAddCommGroup.ofMatrix {M : Matrix n n 𝕜} (hM : M.PosD · simp [h] · exact le_of_lt (hM.re_dotProduct_pos h) definite := fun x (hx : dotProduct _ _ = 0) => by - by_contra' h + by_contra! h simpa [hx, lt_irrefl] using hM.re_dotProduct_pos h add_left := by simp only [star_add, add_dotProduct, eq_self_iff_true, forall_const] smul_left := fun x y r => by diff --git a/Mathlib/LinearAlgebra/Matrix/SchurComplement.lean b/Mathlib/LinearAlgebra/Matrix/SchurComplement.lean index 5276a07eccba1..bef83b3a72173 100644 --- a/Mathlib/LinearAlgebra/Matrix/SchurComplement.lean +++ b/Mathlib/LinearAlgebra/Matrix/SchurComplement.lean @@ -459,16 +459,16 @@ theorem det_add_col_mul_row {A : Matrix m m α} (hA : IsUnit A.det) (u v : m → (A + col u * row v).det = A.det * (1 + row v * A⁻¹ * col u).det := by nth_rewrite 1 [← Matrix.mul_one A] rwa [← Matrix.mul_nonsing_inv_cancel_left A (col u * row v), - ←Matrix.mul_add, det_mul, ←Matrix.mul_assoc, det_one_add_mul_comm, - ←Matrix.mul_assoc] + ← Matrix.mul_add, det_mul, ← Matrix.mul_assoc, det_one_add_mul_comm, + ← Matrix.mul_assoc] /-- A generalization of the **Matrix determinant lemma** -/ theorem det_add_mul {A : Matrix m m α} (U : Matrix m n α) (V : Matrix n m α) (hA : IsUnit A.det) : (A + U * V).det = A.det * (1 + V * A⁻¹ * U).det := by nth_rewrite 1 [← Matrix.mul_one A] - rwa [← Matrix.mul_nonsing_inv_cancel_left A (U * V), ←Matrix.mul_add, - det_mul, ←Matrix.mul_assoc, det_one_add_mul_comm, ←Matrix.mul_assoc] + rwa [← Matrix.mul_nonsing_inv_cancel_left A (U * V), ← Matrix.mul_add, + det_mul, ← Matrix.mul_assoc, det_one_add_mul_comm, ← Matrix.mul_assoc] end Det diff --git a/Mathlib/LinearAlgebra/Matrix/SesquilinearForm.lean b/Mathlib/LinearAlgebra/Matrix/SesquilinearForm.lean index 5a23fb67a843a..74a6c6c02351a 100644 --- a/Mathlib/LinearAlgebra/Matrix/SesquilinearForm.lean +++ b/Mathlib/LinearAlgebra/Matrix/SesquilinearForm.lean @@ -603,9 +603,9 @@ theorem Matrix.isAdjointPair_equiv (P : Matrix n n R) (h : IsUnit P) : simp only [← mul_assoc, P.transpose_nonsing_inv] -- porting note: the previous proof used `conv` and was causing timeouts, so we use `convert` convert this using 2 - · rw [mul_assoc, mul_assoc, ←mul_assoc J] + · rw [mul_assoc, mul_assoc, ← mul_assoc J] rfl - · rw [mul_assoc, mul_assoc, ←mul_assoc _ _ J] + · rw [mul_assoc, mul_assoc, ← mul_assoc _ _ J] rfl rw [Units.eq_mul_inv_iff_mul_eq] conv_rhs => rw [mul_assoc] diff --git a/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean b/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean index 6ade80c39662c..10d07a2e58f01 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean @@ -468,13 +468,13 @@ section Sum @[simp] theorem coeFn_sum {ι : Type*} (Q : ι → QuadraticForm R M) (s : Finset ι) : ⇑(∑ i in s, Q i) = ∑ i in s, ⇑(Q i) := - (coeFnAddMonoidHom : QuadraticForm R M →+ M → R).map_sum Q s + map_sum coeFnAddMonoidHom Q s #align quadratic_form.coe_fn_sum QuadraticForm.coeFn_sum @[simp] theorem sum_apply {ι : Type*} (Q : ι → QuadraticForm R M) (s : Finset ι) (x : M) : (∑ i in s, Q i) x = ∑ i in s, Q i x := - (evalAddMonoidHom x : _ →+ R).map_sum Q s + map_sum (evalAddMonoidHom x : _ →+ R) Q s #align quadratic_form.sum_apply QuadraticForm.sum_apply end Sum @@ -828,7 +828,7 @@ theorem associated_apply (x y : M) : associatedHom S Q x y = ⅟ 2 * (Q (x + y) @[simp] theorem two_nsmul_associated : 2 • associatedHom S Q = Q.polarBilin := by ext dsimp - rw [←smul_mul_assoc, two_nsmul, invOf_two_add_invOf_two, one_mul, polar] + rw [← smul_mul_assoc, two_nsmul, invOf_two_add_invOf_two, one_mul, polar] theorem associated_isSymm : (associatedHom S Q).IsSymm := fun x y => by simp only [associated_apply, add_comm, add_left_comm, sub_eq_add_neg, add_assoc] diff --git a/Mathlib/LinearAlgebra/QuadraticForm/TensorProduct.lean b/Mathlib/LinearAlgebra/QuadraticForm/TensorProduct.lean index 4534063e6217c..53e3d5c799f7d 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/TensorProduct.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/TensorProduct.lean @@ -74,8 +74,8 @@ theorem associated_tmul [Invertible (2 : A)] (Q₁ : QuadraticForm A M₁) (Q₂ theorem polarBilin_tmul [Invertible (2 : A)] (Q₁ : QuadraticForm A M₁) (Q₂ : QuadraticForm R M₂) : polarBilin (Q₁.tmul Q₂) = ⅟(2 : A) • (polarBilin Q₁).tmul (polarBilin Q₂) := by - simp_rw [←two_nsmul_associated A, ←two_nsmul_associated R, BilinForm.tmul, tmul_smul, - ←smul_tmul', map_nsmul, associated_tmul] + simp_rw [← two_nsmul_associated A, ← two_nsmul_associated R, BilinForm.tmul, tmul_smul, + ← smul_tmul', map_nsmul, associated_tmul] rw [smul_comm (_ : A) (_ : ℕ), ← smul_assoc, two_smul _ (_ : A), invOf_two_add_invOf_two, one_smul] @@ -98,7 +98,7 @@ theorem associated_baseChange [Invertible (2 : A)] (Q : QuadraticForm R M₂) : theorem polarBilin_baseChange [Invertible (2 : A)] (Q : QuadraticForm R M₂) : polarBilin (Q.baseChange A) = (polarBilin Q).baseChange A := by rw [QuadraticForm.baseChange, BilinForm.baseChange, polarBilin_tmul, BilinForm.tmul, - ←LinearMap.map_smul, smul_tmul', ←two_nsmul_associated R, coe_associatedHom, associated_sq, + ← LinearMap.map_smul, smul_tmul', ← two_nsmul_associated R, coe_associatedHom, associated_sq, smul_comm, ← smul_assoc, two_smul, invOf_two_add_invOf_two, one_smul] end CommRing diff --git a/Mathlib/LinearAlgebra/StdBasis.lean b/Mathlib/LinearAlgebra/StdBasis.lean index d0c9c0364d155..23df7cef52a10 100644 --- a/Mathlib/LinearAlgebra/StdBasis.lean +++ b/Mathlib/LinearAlgebra/StdBasis.lean @@ -98,7 +98,7 @@ theorem proj_stdBasis_ne (i j : ι) (h : i ≠ j) : (proj i).comp (stdBasis R φ theorem iSup_range_stdBasis_le_iInf_ker_proj (I J : Set ι) (h : Disjoint I J) : ⨆ i ∈ I, range (stdBasis R φ i) ≤ ⨅ i ∈ J, ker (proj i : (∀ i, φ i) →ₗ[R] φ i) := by refine' iSup_le fun i => iSup_le fun hi => range_le_iff_comap.2 _ - simp only [←ker_comp, eq_top_iff, SetLike.le_def, mem_ker, comap_iInf, mem_iInf] + simp only [← ker_comp, eq_top_iff, SetLike.le_def, mem_ker, comap_iInf, mem_iInf] rintro b - j hj rw [proj_stdBasis_ne R φ j i, zero_apply] rintro rfl diff --git a/Mathlib/LinearAlgebra/TensorProduct.lean b/Mathlib/LinearAlgebra/TensorProduct.lean index 38a72fed374d4..8f35bfb38582b 100644 --- a/Mathlib/LinearAlgebra/TensorProduct.lean +++ b/Mathlib/LinearAlgebra/TensorProduct.lean @@ -515,21 +515,8 @@ variable (f : M →ₗ[R] N →ₗ[R] P) with the property that its composition with the canonical bilinear map `M → N → M ⊗ N` is the given bilinear map `M → N → P`. -/ def liftAux : M ⊗[R] N →+ P := - (addConGen (TensorProduct.Eqv R M N)).lift (FreeAddMonoid.lift fun p : M × N => f p.1 p.2) <| - AddCon.addConGen_le fun x y hxy => - match x, y, hxy with - | _, _, Eqv.of_zero_left n => - (AddCon.ker_rel _).2 <| by simp_rw [map_zero, FreeAddMonoid.lift_eval_of, f.map_zero₂] - | _, _, Eqv.of_zero_right m => - (AddCon.ker_rel _).2 <| by simp_rw [map_zero, FreeAddMonoid.lift_eval_of, (f m).map_zero] - | _, _, Eqv.of_add_left m₁ m₂ n => - (AddCon.ker_rel _).2 <| by simp_rw [map_add, FreeAddMonoid.lift_eval_of, f.map_add₂] - | _, _, Eqv.of_add_right m n₁ n₂ => - (AddCon.ker_rel _).2 <| by simp_rw [map_add, FreeAddMonoid.lift_eval_of, (f m).map_add] - | _, _, Eqv.of_smul r m n => - (AddCon.ker_rel _).2 <| by simp_rw [FreeAddMonoid.lift_eval_of, f.map_smul₂, (f m).map_smul] - | _, _, Eqv.add_comm x y => - (AddCon.ker_rel _).2 <| by simp_rw [map_add, add_comm] + liftAddHom (LinearMap.toAddMonoidHom'.comp <| f.toAddMonoidHom) + fun r m n => by dsimp; rw [LinearMap.map_smul₂, map_smul] #align tensor_product.lift_aux TensorProduct.liftAux theorem liftAux_tmul (m n) : liftAux f (m ⊗ₜ n) = f m n := diff --git a/Mathlib/Logic/Basic.lean b/Mathlib/Logic/Basic.lean index 15762c269f6c1..e6e32df4d1e6e 100644 --- a/Mathlib/Logic/Basic.lean +++ b/Mathlib/Logic/Basic.lean @@ -279,6 +279,15 @@ theorem Iff.not_left (h : a ↔ ¬b) : ¬a ↔ b := h.not.trans not_not theorem Iff.not_right (h : ¬a ↔ b) : a ↔ ¬b := not_not.symm.trans h.not #align iff.not_right Iff.not_right +protected lemma Iff.ne {α β : Sort*} {a b : α} {c d : β} : (a = b ↔ c = d) → (a ≠ b ↔ c ≠ d) := + Iff.not + +lemma Iff.ne_left {α β : Sort*} {a b : α} {c d : β} : (a = b ↔ c ≠ d) → (a ≠ b ↔ c = d) := + Iff.not_left + +lemma Iff.ne_right {α β : Sort*} {a b : α} {c d : β} : (a ≠ b ↔ c = d) → (a = b ↔ c ≠ d) := + Iff.not_right + /-! ### Declarations about `Xor'` -/ @[simp] theorem xor_true : Xor' True = Not := by diff --git a/Mathlib/Logic/Embedding/Basic.lean b/Mathlib/Logic/Embedding/Basic.lean index 9b92aa59f028c..317095471c93f 100644 --- a/Mathlib/Logic/Embedding/Basic.lean +++ b/Mathlib/Logic/Embedding/Basic.lean @@ -195,7 +195,7 @@ def setValue {α β} (f : α ↪ β) (a : α) (b : β) [∀ a', Decidable (a' = (try simp only [f.injective.eq_iff, not_true_eq_false] at *) · rw[h₁,h₂] · rw[h₁,h] - · rw[h₅,←h] + · rw[h₅, ← h] · exact h₆.symm · exfalso; exact h₅ h.symm · exfalso; exact h₁ h diff --git a/Mathlib/Logic/Equiv/Defs.lean b/Mathlib/Logic/Equiv/Defs.lean index 36ffcf64bced2..59cbd582509de 100644 --- a/Mathlib/Logic/Equiv/Defs.lean +++ b/Mathlib/Logic/Equiv/Defs.lean @@ -330,6 +330,9 @@ theorem eq_symm_apply {α β} (e : α ≃ β) {x y} : y = e.symm x ↔ e y = x : @[simp] theorem symm_symm (e : α ≃ β) : e.symm.symm = e := by cases e; rfl #align equiv.symm_symm Equiv.symm_symm +theorem symm_bijective : Function.Bijective (Equiv.symm : (α ≃ β) → β ≃ α) := + Function.bijective_iff_has_inverse.mpr ⟨_, symm_symm, symm_symm⟩ + @[simp] theorem trans_refl (e : α ≃ β) : e.trans (Equiv.refl β) = e := by cases e; rfl #align equiv.trans_refl Equiv.trans_refl diff --git a/Mathlib/Logic/Equiv/LocalEquiv.lean b/Mathlib/Logic/Equiv/LocalEquiv.lean index 0a4972be5c260..814b8d495151e 100644 --- a/Mathlib/Logic/Equiv/LocalEquiv.lean +++ b/Mathlib/Logic/Equiv/LocalEquiv.lean @@ -333,6 +333,9 @@ theorem symm_symm : e.symm.symm = e := by rfl #align local_equiv.symm_symm LocalEquiv.symm_symm +theorem symm_bijective : Function.Bijective (LocalEquiv.symm : LocalEquiv α β → LocalEquiv β α) := + Function.bijective_iff_has_inverse.mpr ⟨_, symm_symm, symm_symm⟩ + theorem image_source_eq_target : e '' e.source = e.target := e.bijOn.image_eq #align local_equiv.image_source_eq_target LocalEquiv.image_source_eq_target @@ -677,9 +680,9 @@ protected def trans' (e' : LocalEquiv β γ) (h : e.target = e'.source) : LocalE invFun := e.symm ∘ e'.symm source := e.source target := e'.target - map_source' x hx := by simp [←h, hx] + map_source' x hx := by simp [← h, hx] map_target' y hy := by simp [h, hy] - left_inv' x hx := by simp [hx, ←h] + left_inv' x hx := by simp [hx, ← h] right_inv' y hy := by simp [hy, h] #align local_equiv.trans' LocalEquiv.trans' diff --git a/Mathlib/Logic/Function/Basic.lean b/Mathlib/Logic/Function/Basic.lean index 6132abb5c8313..15998a2906f0d 100644 --- a/Mathlib/Logic/Function/Basic.lean +++ b/Mathlib/Logic/Function/Basic.lean @@ -765,27 +765,11 @@ lemma factorsThrough_iff (g : α → γ) [Nonempty γ] : g.FactorsThrough f ↔ fun h _ _ hf => by rw [Classical.choose_spec h, comp_apply, comp_apply, hf]⟩ #align function.factors_through_iff Function.factorsThrough_iff -lemma FactorsThrough.apply_extend {δ} {g : α → γ} (hf : FactorsThrough g f) - (F : γ → δ) (e' : β → γ) (b : β) : - F (extend f g e' b) = extend f (F ∘ g) (F ∘ e') b := by - by_cases hb : ∃ a, f a = b - case pos => - rcases hb with ⟨a, ha⟩ - subst b - rw [hf.extend_apply, FactorsThrough.extend_apply, comp] - case intro.hf => - intro a b h - simp only [comp_apply] - apply congr_arg - exact hf h - case neg => - rw [extend_apply' _ _ _ hb, extend_apply' _ _ _ hb, comp] -#align function.factors_through.apply_extend Function.FactorsThrough.apply_extend - -lemma Injective.apply_extend {δ} (hf : Injective f) (F : γ → δ) (g : α → γ) (e' : β → γ) (b : β) : +lemma apply_extend {δ} {g : α → γ} (F : γ → δ) (f : α → β) (e' : β → γ) (b : β) : F (extend f g e' b) = extend f (F ∘ g) (F ∘ e') b := - (hf.factorsThrough g).apply_extend F e' b -#align function.injective.apply_extend Function.Injective.apply_extend + apply_dite F _ _ _ +#align function.factors_through.apply_extend Function.apply_extend +#align function.injective.apply_extend Function.apply_extend theorem extend_injective (hf : Injective f) (e' : β → γ) : Injective fun g ↦ extend f g e' := by intro g₁ g₂ hg diff --git a/Mathlib/Logic/Function/Iterate.lean b/Mathlib/Logic/Function/Iterate.lean index c2739886ced8e..09f3b151eb3e6 100644 --- a/Mathlib/Logic/Function/Iterate.lean +++ b/Mathlib/Logic/Function/Iterate.lean @@ -236,7 +236,7 @@ theorem iterate_commute (m n : ℕ) : Commute (fun f : α → α ↦ f^[m]) fun #align function.iterate_commute Function.iterate_commute lemma iterate_add_eq_iterate (hf : Injective f) : f^[m + n] a = f^[n] a ↔ f^[m] a = a := - Iff.trans (by rw [←iterate_add_apply, Nat.add_comm]) (hf.iterate n).eq_iff + Iff.trans (by rw [← iterate_add_apply, Nat.add_comm]) (hf.iterate n).eq_iff #align function.iterate_add_eq_iterate Function.iterate_add_eq_iterate alias ⟨iterate_cancel_of_add, _⟩ := iterate_add_eq_iterate diff --git a/Mathlib/Logic/Relation.lean b/Mathlib/Logic/Relation.lean index d300b70d8faa5..b07b462feff16 100644 --- a/Mathlib/Logic/Relation.lean +++ b/Mathlib/Logic/Relation.lean @@ -3,6 +3,7 @@ Copyright (c) 2018 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl -/ +import Mathlib.Logic.Function.Basic import Mathlib.Logic.Relator import Mathlib.Init.Propext import Mathlib.Init.Data.Quot @@ -47,7 +48,7 @@ the bundled version, see `Rel`. open Function -variable {α β γ δ : Type*} +variable {α β γ δ ε ζ : Type*} section NeImp @@ -206,6 +207,9 @@ theorem _root_.Acc.of_downward_closed (dc : ∀ {a b}, rβ b (f a) → ∃ c, f end Fibration +section map +variable {r : α → β → Prop} {f : α → γ} {g : β → δ} {c : γ} {d : δ} + /-- The map of a relation `r` through a pair of functions pushes the relation to the codomains of the functions. The resulting relation is defined by having pairs of terms related if they have preimages @@ -215,6 +219,22 @@ protected def Map (r : α → β → Prop) (f : α → γ) (g : β → δ) : γ ∃ a b, r a b ∧ f a = c ∧ g b = d #align relation.map Relation.Map +@[simp] lemma map_map (r : α → β → Prop) (f₁ : α → γ) (g₁ : β → δ) (f₂ : γ → ε) (g₂ : δ → ζ) : + Relation.Map (Relation.Map r f₁ g₁) f₂ g₂ = Relation.Map r (f₂ ∘ f₁) (g₂ ∘ g₁) := by + ext a b + simp_rw [Relation.Map, Function.comp_apply, ←exists_and_right, @exists_comm γ, @exists_comm δ] + refine' exists₂_congr fun a b ↦ ⟨_, fun h ↦ ⟨_, _, ⟨⟨h.1, rfl, rfl⟩, h.2⟩⟩⟩ + rintro ⟨_, _, ⟨hab, rfl, rfl⟩, h⟩ + exact ⟨hab, h⟩ + +@[simp] +lemma map_apply_apply (hf : Injective f) (hg : Injective g) (r : α → β → Prop) (a : α) (b : β) : + Relation.Map r f g (f a) (g b) ↔ r a b := by simp [Relation.Map, hf.eq_iff, hg.eq_iff] + +@[simp] lemma map_id_id (r : α → β → Prop) : Relation.Map r id id = r := by ext; simp [Relation.Map] + +end map + variable {r : α → α → Prop} {a b c d : α} /-- `ReflTransGen r`: reflexive transitive closure of `r` -/ @@ -615,10 +635,10 @@ theorem reflTransGen_swap : ReflTransGen (swap r) a b ↔ ReflTransGen r b a := · exact TransGen.mono (fun _ _ ↦ .single) h @[simp] lemma reflTransGen_reflGen : ReflTransGen (ReflGen r) = ReflTransGen r := by - simp only [←transGen_reflGen, reflGen_eq_self reflexive_reflGen] + simp only [← transGen_reflGen, reflGen_eq_self reflexive_reflGen] @[simp] lemma reflTransGen_transGen : ReflTransGen (TransGen r) = ReflTransGen r := by - simp only [←reflGen_transGen, transGen_idem] + simp only [← reflGen_transGen, transGen_idem] lemma reflTransGen_eq_transGen (hr : Reflexive r) : ReflTransGen r = TransGen r := by diff --git a/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean b/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean index 73de6ff1a4a1e..7f3e324a8a0c1 100644 --- a/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean +++ b/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean @@ -105,6 +105,9 @@ theorem isPiSystem_isOpen [TopologicalSpace α] : IsPiSystem (IsOpen : Set α fun _s hs _t ht _ => IsOpen.inter hs ht #align is_pi_system_is_open isPiSystem_isOpen +lemma isPiSystem_isClosed [TopologicalSpace α] : IsPiSystem (IsClosed : Set α → Prop) := + fun _s hs _t ht _ ↦ IsClosed.inter hs ht + theorem borel_eq_generateFrom_isClosed [TopologicalSpace α] : borel α = .generateFrom { s | IsClosed s } := le_antisymm diff --git a/Mathlib/MeasureTheory/Constructions/Pi.lean b/Mathlib/MeasureTheory/Constructions/Pi.lean index cf4df36a44dd1..3ad72609dec61 100644 --- a/Mathlib/MeasureTheory/Constructions/Pi.lean +++ b/Mathlib/MeasureTheory/Constructions/Pi.lean @@ -243,7 +243,7 @@ instance sigmaFinite_tprod (l : List δ) (μ : ∀ i, Measure (π i)) [∀ i, Si SigmaFinite (Measure.tprod l μ) := by induction l with | nil => rw [tprod_nil]; infer_instance - | cons i l ih => rw [tprod_cons]; exact @prod.instSigmaFinite _ _ _ _ _ _ ih _ + | cons i l ih => rw [tprod_cons]; exact @prod.instSigmaFinite _ _ _ _ _ _ _ ih #align measure_theory.measure.sigma_finite_tprod MeasureTheory.Measure.sigmaFinite_tprod theorem tprod_tprod (l : List δ) (μ : ∀ i, Measure (π i)) [∀ i, SigmaFinite (μ i)] diff --git a/Mathlib/MeasureTheory/Constructions/Polish.lean b/Mathlib/MeasureTheory/Constructions/Polish.lean index 1f8bb4439a808..5e818e3b2203b 100644 --- a/Mathlib/MeasureTheory/Constructions/Polish.lean +++ b/Mathlib/MeasureTheory/Constructions/Polish.lean @@ -749,7 +749,7 @@ theorem measurableSet_range_of_continuous_injective {β : Type*} [TopologicalSpa have I : ∀ m n, ((s m).1 ∩ (s n).1).Nonempty := by intro m n rw [← not_disjoint_iff_nonempty_inter] - by_contra' h + by_contra! h have A : x ∈ q ⟨(s m, s n), h⟩ \ q ⟨(s n, s m), h.symm⟩ := haveI := mem_iInter.1 (hxs m).2 (s n) (mem_iInter.1 this h : _) @@ -777,7 +777,7 @@ theorem measurableSet_range_of_continuous_injective {β : Type*} [TopologicalSpa rw [← this] exact mem_range_self _ -- assume for a contradiction that `f z ≠ x`. - by_contra' hne + by_contra! hne -- introduce disjoint open sets `v` and `w` separating `f z` from `x`. obtain ⟨v, w, v_open, w_open, fzv, xw, hvw⟩ := t2_separation hne obtain ⟨δ, δpos, hδ⟩ : ∃ δ > (0 : ℝ), ball z δ ⊆ f ⁻¹' v := by diff --git a/Mathlib/MeasureTheory/Constructions/Prod/Basic.lean b/Mathlib/MeasureTheory/Constructions/Prod/Basic.lean index 458bd94ce2b92..f4f2064acf9d4 100644 --- a/Mathlib/MeasureTheory/Constructions/Prod/Basic.lean +++ b/Mathlib/MeasureTheory/Constructions/Prod/Basic.lean @@ -14,7 +14,7 @@ import Mathlib.MeasureTheory.Measure.OpenPos # The product measure In this file we define and prove properties about the binary product measure. If `α` and `β` have -σ-finite measures `μ` resp. `ν` then `α × β` can be equipped with a σ-finite measure `μ.prod ν` that +s-finite measures `μ` resp. `ν` then `α × β` can be equipped with a s-finite measure `μ.prod ν` that satisfies `(μ.prod ν) s = ∫⁻ x, ν {y | (x, y) ∈ s} ∂μ`. We also have `(μ.prod ν) (s ×ˢ t) = μ s * ν t`, i.e. the measure of a rectangle is the product of the measures of the sides. @@ -174,32 +174,30 @@ theorem measurable_measure_prod_mk_left_finite [IsFiniteMeasure ν] {s : Set (α apply Measurable.ennreal_tsum h3f #align measurable_measure_prod_mk_left_finite measurable_measure_prod_mk_left_finite -/-- If `ν` is a σ-finite measure, and `s ⊆ α × β` is measurable, then `x ↦ ν { y | (x, y) ∈ s }` is - a measurable function. -/ -theorem measurable_measure_prod_mk_left [SigmaFinite ν] {s : Set (α × β)} (hs : MeasurableSet s) : +/-- If `ν` is an s-finite measure, and `s ⊆ α × β` is measurable, then `x ↦ ν { y | (x, y) ∈ s }` + is a measurable function. -/ +theorem measurable_measure_prod_mk_left [SFinite ν] {s : Set (α × β)} (hs : MeasurableSet s) : Measurable fun x => ν (Prod.mk x ⁻¹' s) := by - have : ∀ x, MeasurableSet (Prod.mk x ⁻¹' s) := fun x => measurable_prod_mk_left hs - simp only [← @iSup_restrict_spanningSets _ _ ν, this] - apply measurable_iSup; intro i - haveI := Fact.mk (measure_spanningSets_lt_top ν i) - exact measurable_measure_prod_mk_left_finite hs + rw [← sum_sFiniteSeq ν] + simp_rw [Measure.sum_apply_of_countable] + exact Measurable.ennreal_tsum (fun i ↦ measurable_measure_prod_mk_left_finite hs) #align measurable_measure_prod_mk_left measurable_measure_prod_mk_left /-- If `μ` is a σ-finite measure, and `s ⊆ α × β` is measurable, then `y ↦ μ { x | (x, y) ∈ s }` is a measurable function. -/ -theorem measurable_measure_prod_mk_right {μ : Measure α} [SigmaFinite μ] {s : Set (α × β)} +theorem measurable_measure_prod_mk_right {μ : Measure α} [SFinite μ] {s : Set (α × β)} (hs : MeasurableSet s) : Measurable fun y => μ ((fun x => (x, y)) ⁻¹' s) := measurable_measure_prod_mk_left (measurableSet_swap_iff.mpr hs) #align measurable_measure_prod_mk_right measurable_measure_prod_mk_right -theorem Measurable.map_prod_mk_left [SigmaFinite ν] : +theorem Measurable.map_prod_mk_left [SFinite ν] : Measurable fun x : α => map (Prod.mk x) ν := by apply measurable_of_measurable_coe; intro s hs simp_rw [map_apply measurable_prod_mk_left hs] exact measurable_measure_prod_mk_left hs #align measurable.map_prod_mk_left Measurable.map_prod_mk_left -theorem Measurable.map_prod_mk_right {μ : Measure α} [SigmaFinite μ] : +theorem Measurable.map_prod_mk_right {μ : Measure α} [SFinite μ] : Measurable fun y : β => map (fun x : α => (x, y)) μ := by apply measurable_of_measurable_coe; intro s hs simp_rw [map_apply measurable_prod_mk_right hs] @@ -237,7 +235,7 @@ theorem MeasurableEmbedding.prod_mk {α β γ δ : Type*} {mα : MeasurableSpace /-- The Lebesgue integral is measurable. This shows that the integrand of (the right-hand-side of) Tonelli's theorem is measurable. -/ -theorem Measurable.lintegral_prod_right' [SigmaFinite ν] : +theorem Measurable.lintegral_prod_right' [SFinite ν] : ∀ {f : α × β → ℝ≥0∞}, Measurable f → Measurable fun x => ∫⁻ y, f (x, y) ∂ν := by have m := @measurable_prod_mk_left refine' Measurable.ennreal_induction (P := fun f => Measurable fun (x : α) => ∫⁻ y, f (x, y) ∂ν) @@ -260,14 +258,14 @@ theorem Measurable.lintegral_prod_right' [SigmaFinite ν] : /-- The Lebesgue integral is measurable. This shows that the integrand of (the right-hand-side of) Tonelli's theorem is measurable. This version has the argument `f` in curried form. -/ -theorem Measurable.lintegral_prod_right [SigmaFinite ν] {f : α → β → ℝ≥0∞} +theorem Measurable.lintegral_prod_right [SFinite ν] {f : α → β → ℝ≥0∞} (hf : Measurable (uncurry f)) : Measurable fun x => ∫⁻ y, f x y ∂ν := hf.lintegral_prod_right' #align measurable.lintegral_prod_right Measurable.lintegral_prod_right /-- The Lebesgue integral is measurable. This shows that the integrand of (the right-hand-side of) the symmetric version of Tonelli's theorem is measurable. -/ -theorem Measurable.lintegral_prod_left' [SigmaFinite μ] {f : α × β → ℝ≥0∞} (hf : Measurable f) : +theorem Measurable.lintegral_prod_left' [SFinite μ] {f : α × β → ℝ≥0∞} (hf : Measurable f) : Measurable fun y => ∫⁻ x, f (x, y) ∂μ := (measurable_swap_iff.mpr hf).lintegral_prod_right' #align measurable.lintegral_prod_left' Measurable.lintegral_prod_left' @@ -275,7 +273,7 @@ theorem Measurable.lintegral_prod_left' [SigmaFinite μ] {f : α × β → ℝ /-- The Lebesgue integral is measurable. This shows that the integrand of (the right-hand-side of) the symmetric version of Tonelli's theorem is measurable. This version has the argument `f` in curried form. -/ -theorem Measurable.lintegral_prod_left [SigmaFinite μ] {f : α → β → ℝ≥0∞} +theorem Measurable.lintegral_prod_left [SFinite μ] {f : α → β → ℝ≥0∞} (hf : Measurable (uncurry f)) : Measurable fun y => ∫⁻ x, f x y ∂μ := hf.lintegral_prod_left' #align measurable.lintegral_prod_left Measurable.lintegral_prod_left @@ -288,7 +286,7 @@ namespace MeasureTheory namespace Measure /-- The binary product of measures. They are defined for arbitrary measures, but we basically - prove all properties under the assumption that at least one of them is σ-finite. -/ + prove all properties under the assumption that at least one of them is s-finite. -/ protected irreducible_def prod (μ : Measure α) (ν : Measure β) : Measure (α × β) := bind μ fun x : α => map (Prod.mk x) ν #align measure_theory.measure.prod MeasureTheory.Measure.prod @@ -302,7 +300,7 @@ theorem volume_eq_prod (α β) [MeasureSpace α] [MeasureSpace β] : rfl #align measure_theory.measure.volume_eq_prod MeasureTheory.Measure.volume_eq_prod -variable [SigmaFinite ν] +variable [SFinite ν] theorem prod_apply {s : Set (α × β)} (hs : MeasurableSet s) : μ.prod ν s = ∫⁻ x, ν (Prod.mk x ⁻¹' s) ∂μ := by @@ -354,7 +352,7 @@ theorem prod_prod (s : Set α) (t : Set β) : μ.prod ν (s ×ˢ t) = μ s * ν instance prod.instIsOpenPosMeasure {X Y : Type*} [TopologicalSpace X] [TopologicalSpace Y] {m : MeasurableSpace X} {μ : Measure X} [IsOpenPosMeasure μ] {m' : MeasurableSpace Y} - {ν : Measure Y} [IsOpenPosMeasure ν] [SigmaFinite ν] : IsOpenPosMeasure (μ.prod ν) := by + {ν : Measure Y} [IsOpenPosMeasure ν] [SFinite ν] : IsOpenPosMeasure (μ.prod ν) := by constructor rintro U U_open ⟨⟨x, y⟩, hxy⟩ rcases isOpen_prod_iff.1 U_open x y hxy with ⟨u, v, u_open, v_open, xu, yv, huv⟩ @@ -368,7 +366,7 @@ instance prod.instIsOpenPosMeasure {X Y : Type*} [TopologicalSpace X] [Topologic instance {X Y : Type*} [TopologicalSpace X] [MeasureSpace X] [IsOpenPosMeasure (volume : Measure X)] [TopologicalSpace Y] [MeasureSpace Y] [IsOpenPosMeasure (volume : Measure Y)] - [SigmaFinite (volume : Measure Y)] : IsOpenPosMeasure (volume : Measure (X × Y)) := + [SFinite (volume : Measure Y)] : IsOpenPosMeasure (volume : Measure (X × Y)) := prod.instIsOpenPosMeasure instance prod.instIsFiniteMeasure {α β : Type*} {mα : MeasurableSpace α} {mβ : MeasurableSpace β} @@ -396,7 +394,7 @@ instance {α β : Type*} [MeasureSpace α] [MeasureSpace β] instance prod.instIsFiniteMeasureOnCompacts {α β : Type*} [TopologicalSpace α] [TopologicalSpace β] {mα : MeasurableSpace α} {mβ : MeasurableSpace β} (μ : Measure α) (ν : Measure β) - [IsFiniteMeasureOnCompacts μ] [IsFiniteMeasureOnCompacts ν] [SigmaFinite ν] : + [IsFiniteMeasureOnCompacts μ] [IsFiniteMeasureOnCompacts ν] [SFinite ν] : IsFiniteMeasureOnCompacts (μ.prod ν) := by refine' ⟨fun K hK => _⟩ set L := (Prod.fst '' K) ×ˢ (Prod.snd '' K) with hL @@ -414,7 +412,7 @@ instance prod.instIsFiniteMeasureOnCompacts {α β : Type*} [TopologicalSpace α instance {X Y : Type*} [TopologicalSpace X] [MeasureSpace X] [IsFiniteMeasureOnCompacts (volume : Measure X)] [TopologicalSpace Y] [MeasureSpace Y] [IsFiniteMeasureOnCompacts (volume : Measure Y)] - [SigmaFinite (volume : Measure Y)] : IsFiniteMeasureOnCompacts (volume : Measure (X × Y)) := + [SFinite (volume : Measure Y)] : IsFiniteMeasureOnCompacts (volume : Measure (X × Y)) := prod.instIsFiniteMeasureOnCompacts _ _ instance prod.instNoAtoms_fst [NoAtoms μ] : @@ -453,7 +451,7 @@ theorem measure_ae_null_of_prod_null {s : Set (α × β)} (h : μ.prod ν s = 0) eventually_of_forall fun x => zero_le _⟩ #align measure_theory.measure.measure_ae_null_of_prod_null MeasureTheory.Measure.measure_ae_null_of_prod_null -theorem AbsolutelyContinuous.prod [SigmaFinite ν'] (h1 : μ ≪ μ') (h2 : ν ≪ ν') : +theorem AbsolutelyContinuous.prod [SFinite ν'] (h1 : μ ≪ μ') (h2 : ν ≪ ν') : μ.prod ν ≪ μ'.prod ν' := by refine' AbsolutelyContinuous.mk fun s hs h2s => _ rw [measure_prod_null hs] at h2s ⊢ @@ -546,16 +544,52 @@ noncomputable def FiniteSpanningSetsIn.prod {ν : Measure β} {C : Set (Set α)} · simp_rw [iUnion_unpair_prod, hμ.spanning, hν.spanning, univ_prod_univ] #align measure_theory.measure.finite_spanning_sets_in.prod MeasureTheory.Measure.FiniteSpanningSetsIn.prod -variable [SigmaFinite μ] +lemma prod_sum_left {ι : Type*} (m : ι → Measure α) (μ : Measure β) [SFinite μ] : + (Measure.sum m).prod μ = Measure.sum (fun i ↦ (m i).prod μ) := by + ext s hs + simp only [prod_apply hs, lintegral_sum_measure, hs, sum_apply, ENNReal.tsum_prod'] +#align measure_theory.measure.sum_prod MeasureTheory.Measure.prod_sum_left -instance prod.instSigmaFinite : SigmaFinite (μ.prod ν) := +lemma prod_sum_right {ι' : Type*} [Countable ι'] (m : Measure α) (m' : ι' → Measure β) + [∀ n, SFinite (m' n)] : + m.prod (Measure.sum m') = Measure.sum (fun p ↦ m.prod (m' p)) := by + ext s hs + simp only [prod_apply hs, lintegral_sum_measure, hs, sum_apply, ENNReal.tsum_prod'] + have M : ∀ x, MeasurableSet (Prod.mk x ⁻¹' s) := fun x => measurable_prod_mk_left hs + simp_rw [Measure.sum_apply _ (M _)] + rw [lintegral_tsum (fun i ↦ (measurable_measure_prod_mk_left hs).aemeasurable)] +#align measure_theory.measure.prod_sum MeasureTheory.Measure.prod_sum_right + +lemma prod_sum {ι ι' : Type*} [Countable ι'] (m : ι → Measure α) (m' : ι' → Measure β) + [∀ n, SFinite (m' n)] : + (Measure.sum m).prod (Measure.sum m') = + Measure.sum (fun (p : ι × ι') ↦ (m p.1).prod (m' p.2)) := by + simp_rw [prod_sum_left, prod_sum_right, sum_sum] + +instance prod.instSigmaFinite {α β : Type*} {_ : MeasurableSpace α} {μ : Measure α} + [SigmaFinite μ] {_ : MeasurableSpace β} {ν : Measure β} [SigmaFinite ν] : + SigmaFinite (μ.prod ν) := (μ.toFiniteSpanningSetsIn.prod ν.toFiniteSpanningSetsIn).sigmaFinite #align measure_theory.measure.prod.sigma_finite MeasureTheory.Measure.prod.instSigmaFinite +instance prod.instSFinite {α β : Type*} {_ : MeasurableSpace α} {μ : Measure α} + [SFinite μ] {_ : MeasurableSpace β} {ν : Measure β} [SFinite ν] : + SFinite (μ.prod ν) := by + have : μ.prod ν = + Measure.sum (fun (p : ℕ × ℕ) ↦ (sFiniteSeq μ p.1).prod (sFiniteSeq ν p.2)) := by + conv_lhs => rw [← sum_sFiniteSeq μ, ← sum_sFiniteSeq ν] + apply prod_sum + rw [this] + infer_instance + instance {α β} [MeasureSpace α] [SigmaFinite (volume : Measure α)] [MeasureSpace β] [SigmaFinite (volume : Measure β)] : SigmaFinite (volume : Measure (α × β)) := prod.instSigmaFinite +instance {α β} [MeasureSpace α] [SFinite (volume : Measure α)] + [MeasureSpace β] [SFinite (volume : Measure β)] : SFinite (volume : Measure (α × β)) := + prod.instSFinite + /-- A measure on a product space equals the product measure if they are equal on rectangles with as sides sets that generate the corresponding σ-algebras. -/ theorem prod_eq_generateFrom {μ : Measure α} {ν : Measure β} {C : Set (Set α)} {D : Set (Set β)} @@ -571,15 +605,35 @@ theorem prod_eq_generateFrom {μ : Measure α} {ν : Measure β} {C : Set (Set rw [h₁ s hs t ht, prod_prod] #align measure_theory.measure.prod_eq_generate_from MeasureTheory.Measure.prod_eq_generateFrom -/-- A measure on a product space equals the product measure if they are equal on rectangles. -/ -theorem prod_eq {μν : Measure (α × β)} +/- Note that the next theorem is not true for s-finite measures: let `μ = ν = ∞ • Leb` on `[0,1]` +(they are s-finite as countable sums of the finite Lebesgue measure), and let `μν = μ.prod ν + λ` +where `λ` is Lebesgue measure on the diagonal. Then both measures give infinite mass to rectangles +`s × t` whose sides have positive Lebesgue measure, and `0` measure when one of the sides has zero +Lebesgue measure. And yet they do not coincide, as the first one gives zero mass to the diagonal, +and the second one gives mass one. +-/ +/-- A measure on a product space equals the product measure of sigma-finite measures if they are +equal on rectangles. -/ +theorem prod_eq {μ : Measure α} [SigmaFinite μ] {ν : Measure β} [SigmaFinite ν] + {μν : Measure (α × β)} (h : ∀ s t, MeasurableSet s → MeasurableSet t → μν (s ×ˢ t) = μ s * ν t) : μ.prod ν = μν := prod_eq_generateFrom generateFrom_measurableSet generateFrom_measurableSet isPiSystem_measurableSet isPiSystem_measurableSet μ.toFiniteSpanningSetsIn ν.toFiniteSpanningSetsIn fun s hs t ht => h s t hs ht #align measure_theory.measure.prod_eq MeasureTheory.Measure.prod_eq +variable [SFinite μ] + theorem prod_swap : map Prod.swap (μ.prod ν) = ν.prod μ := by + have : sum (fun (i : ℕ × ℕ) ↦ map Prod.swap ((sFiniteSeq μ i.1).prod (sFiniteSeq ν i.2))) + = sum (fun (i : ℕ × ℕ) ↦ map Prod.swap ((sFiniteSeq μ i.2).prod (sFiniteSeq ν i.1))) := by + ext s hs + rw [sum_apply _ hs, sum_apply _ hs] + exact ((Equiv.prodComm ℕ ℕ).tsum_eq _).symm + rw [← sum_sFiniteSeq μ, ← sum_sFiniteSeq ν, prod_sum, prod_sum, + map_sum measurable_swap.aemeasurable, this] + congr 1 + ext1 i refine' (prod_eq _).symm intro s t hs ht simp_rw [map_apply measurable_swap (hs.prod ht), preimage_swap_prod, prod_prod, mul_comm] @@ -632,11 +686,23 @@ lemma nullMeasurableSet_prod {s : Set α} {t : Set β} : rcases eq_or_ne (ν t) 0 with ht | ht; · simp [NullMeasurableSet.of_null, *] simp [*, nullMeasurableSet_prod_of_ne_zero] -theorem prodAssoc_prod [SigmaFinite τ] : +theorem prodAssoc_prod [SFinite τ] : map MeasurableEquiv.prodAssoc ((μ.prod ν).prod τ) = μ.prod (ν.prod τ) := by + have : sum (fun (p : ℕ × ℕ × ℕ) ↦ + (sFiniteSeq μ p.1).prod ((sFiniteSeq ν p.2.1).prod (sFiniteSeq τ p.2.2))) + = sum (fun (p : (ℕ × ℕ) × ℕ) ↦ + (sFiniteSeq μ p.1.1).prod ((sFiniteSeq ν p.1.2).prod (sFiniteSeq τ p.2))) := by + ext s hs + rw [sum_apply _ hs, sum_apply _ hs, ← (Equiv.prodAssoc _ _ _).tsum_eq] + rfl + rw [← sum_sFiniteSeq μ, ← sum_sFiniteSeq ν, ← sum_sFiniteSeq τ, prod_sum, prod_sum, + map_sum MeasurableEquiv.prodAssoc.measurable.aemeasurable, prod_sum, prod_sum, this] + congr + ext1 i refine' (prod_eq_generateFrom generateFrom_measurableSet generateFrom_prod - isPiSystem_measurableSet isPiSystem_prod μ.toFiniteSpanningSetsIn - (ν.toFiniteSpanningSetsIn.prod τ.toFiniteSpanningSetsIn) _).symm + isPiSystem_measurableSet isPiSystem_prod ((sFiniteSeq μ i.1.1)).toFiniteSpanningSetsIn + ((sFiniteSeq ν i.1.2).toFiniteSpanningSetsIn.prod (sFiniteSeq τ i.2).toFiniteSpanningSetsIn) + _).symm rintro s hs _ ⟨t, u, ht, hu, rfl⟩; rw [mem_setOf_eq] at hs ht hu simp_rw [map_apply (MeasurableEquiv.measurable _) (hs.prod (ht.prod hu)), MeasurableEquiv.prodAssoc, MeasurableEquiv.coe_mk, Equiv.prod_assoc_preimage, prod_prod, @@ -647,6 +713,10 @@ theorem prodAssoc_prod [SigmaFinite τ] : theorem prod_restrict (s : Set α) (t : Set β) : (μ.restrict s).prod (ν.restrict t) = (μ.prod ν).restrict (s ×ˢ t) := by + rw [← sum_sFiniteSeq μ, ← sum_sFiniteSeq ν, restrict_sum_of_countable, restrict_sum_of_countable, + prod_sum, prod_sum, restrict_sum_of_countable] + congr 1 + ext1 i refine' prod_eq fun s' t' hs' ht' => _ rw [restrict_apply (hs'.prod ht'), prod_inter_prod, prod_prod, restrict_apply hs', restrict_apply ht'] @@ -659,39 +729,41 @@ theorem restrict_prod_eq_prod_univ (s : Set α) : #align measure_theory.measure.restrict_prod_eq_prod_univ MeasureTheory.Measure.restrict_prod_eq_prod_univ theorem prod_dirac (y : β) : μ.prod (dirac y) = map (fun x => (x, y)) μ := by + rw [← sum_sFiniteSeq μ, prod_sum_left, map_sum measurable_prod_mk_right.aemeasurable] + congr + ext1 i refine' prod_eq fun s t hs ht => _ simp_rw [map_apply measurable_prod_mk_right (hs.prod ht), mk_preimage_prod_left_eq_if, measure_if, - dirac_apply' _ ht, ← indicator_mul_right _ fun _ => μ s, Pi.one_apply, mul_one] + dirac_apply' _ ht, ← indicator_mul_right _ fun _ => sFiniteSeq μ i s, Pi.one_apply, mul_one] #align measure_theory.measure.prod_dirac MeasureTheory.Measure.prod_dirac theorem dirac_prod (x : α) : (dirac x).prod ν = map (Prod.mk x) ν := by + rw [← sum_sFiniteSeq ν, prod_sum_right, map_sum measurable_prod_mk_left.aemeasurable] + congr + ext1 i refine' prod_eq fun s t hs ht => _ simp_rw [map_apply measurable_prod_mk_left (hs.prod ht), mk_preimage_prod_right_eq_if, measure_if, - dirac_apply' _ hs, ← indicator_mul_left _ _ fun _ => ν t, Pi.one_apply, one_mul] + dirac_apply' _ hs, ← indicator_mul_left _ _ fun _ => sFiniteSeq ν i t, Pi.one_apply, one_mul] #align measure_theory.measure.dirac_prod MeasureTheory.Measure.dirac_prod theorem dirac_prod_dirac {x : α} {y : β} : (dirac x).prod (dirac y) = dirac (x, y) := by rw [prod_dirac, map_dirac measurable_prod_mk_right] #align measure_theory.measure.dirac_prod_dirac MeasureTheory.Measure.dirac_prod_dirac -theorem prod_sum {ι : Type*} [Finite ι] (ν : ι → Measure β) [∀ i, SigmaFinite (ν i)] : - μ.prod (sum ν) = sum fun i => μ.prod (ν i) := by - refine' prod_eq fun s t hs ht => _ - simp_rw [sum_apply _ (hs.prod ht), sum_apply _ ht, prod_prod, ENNReal.tsum_mul_left] -#align measure_theory.measure.prod_sum MeasureTheory.Measure.prod_sum - -theorem sum_prod {ι : Type*} [Finite ι] (μ : ι → Measure α) [∀ i, SigmaFinite (μ i)] : - (sum μ).prod ν = sum fun i => (μ i).prod ν := by - refine' prod_eq fun s t hs ht => _ - simp_rw [sum_apply _ (hs.prod ht), sum_apply _ hs, prod_prod, ENNReal.tsum_mul_right] -#align measure_theory.measure.sum_prod MeasureTheory.Measure.sum_prod - -theorem prod_add (ν' : Measure β) [SigmaFinite ν'] : μ.prod (ν + ν') = μ.prod ν + μ.prod ν' := by +theorem prod_add (ν' : Measure β) [SFinite ν'] : μ.prod (ν + ν') = μ.prod ν + μ.prod ν' := by + simp_rw [← sum_sFiniteSeq ν, ← sum_sFiniteSeq ν', sum_add_sum, ← sum_sFiniteSeq μ, prod_sum, + sum_add_sum] + congr + ext1 i refine' prod_eq fun s t _ _ => _ simp_rw [add_apply, prod_prod, left_distrib] #align measure_theory.measure.prod_add MeasureTheory.Measure.prod_add -theorem add_prod (μ' : Measure α) [SigmaFinite μ'] : (μ + μ').prod ν = μ.prod ν + μ'.prod ν := by +theorem add_prod (μ' : Measure α) [SFinite μ'] : (μ + μ').prod ν = μ.prod ν + μ'.prod ν := by + simp_rw [← sum_sFiniteSeq μ, ← sum_sFiniteSeq μ', sum_add_sum, ← sum_sFiniteSeq ν, prod_sum, + sum_add_sum] + congr + ext1 i refine' prod_eq fun s t _ _ => _ simp_rw [add_apply, prod_prod, right_distrib] #align measure_theory.measure.add_prod MeasureTheory.Measure.add_prod @@ -706,11 +778,13 @@ theorem zero_prod (ν : Measure β) : (0 : Measure α).prod ν = 0 := by theorem prod_zero (μ : Measure α) : μ.prod (0 : Measure β) = 0 := by simp [Measure.prod] #align measure_theory.measure.prod_zero MeasureTheory.Measure.prod_zero -theorem map_prod_map {δ} [MeasurableSpace δ] {f : α → β} {g : γ → δ} {μa : Measure α} - {μc : Measure γ} (hfa : SigmaFinite (map f μa)) (hgc : SigmaFinite (map g μc)) - (hf : Measurable f) (hg : Measurable g) : +theorem map_prod_map {δ} [MeasurableSpace δ] {f : α → β} {g : γ → δ} (μa : Measure α) + (μc : Measure γ) [SFinite μa] [SFinite μc] (hf : Measurable f) (hg : Measurable g) : (map f μa).prod (map g μc) = map (Prod.map f g) (μa.prod μc) := by - haveI := hgc.of_map μc hg.aemeasurable + simp_rw [← sum_sFiniteSeq μa, ← sum_sFiniteSeq μc, map_sum hf.aemeasurable, + map_sum hg.aemeasurable, prod_sum, map_sum (hf.prod_map hg).aemeasurable] + congr + ext1 i refine' prod_eq fun s t hs ht => _ rw [map_apply (hf.prod_map hg) (hs.prod ht), map_apply hf hs, map_apply hg ht] exact prod_prod (f ⁻¹' s) (g ⁻¹' t) @@ -725,38 +799,34 @@ namespace MeasurePreserving variable {δ : Type*} [MeasurableSpace δ] {μa : Measure α} {μb : Measure β} {μc : Measure γ} {μd : Measure δ} -theorem skew_product [SigmaFinite μb] [SigmaFinite μd] {f : α → β} (hf : MeasurePreserving f μa μb) +theorem skew_product [SFinite μa] [SFinite μc] {f : α → β} (hf : MeasurePreserving f μa μb) {g : α → γ → δ} (hgm : Measurable (uncurry g)) (hg : ∀ᵐ x ∂μa, map (g x) μc = μd) : MeasurePreserving (fun p : α × γ => (f p.1, g p.1 p.2)) (μa.prod μc) (μb.prod μd) := by classical - have : Measurable fun p : α × γ => (f p.1, g p.1 p.2) := (hf.1.comp measurable_fst).prod_mk hgm - /- if `μa = 0`, then the lemma is trivial, otherwise we can use `hg` - to deduce `SigmaFinite μc`. -/ - rcases eq_or_ne μa 0 with (rfl | ha) - · rw [← hf.map_eq, zero_prod, Measure.map_zero, zero_prod] - exact ⟨this, by simp only [Measure.map_zero]⟩ - have sf : SigmaFinite μc := by - rcases (ae_neBot.2 ha).nonempty_of_mem hg with ⟨x, hx : map (g x) μc = μd⟩ - exact SigmaFinite.of_map _ hgm.of_uncurry_left.aemeasurable (by rwa [hx]) - -- Thus we can apply `Measure.prod_eq` to prove equality of measures. - refine' ⟨this, (prod_eq fun s t hs ht => _).symm⟩ - rw [map_apply this (hs.prod ht)] - refine' (prod_apply (this <| hs.prod ht)).trans _ - have : ∀ᵐ x ∂μa, - μc ((fun y => (f x, g x y)) ⁻¹' s ×ˢ t) = indicator (f ⁻¹' s) (fun _ => μd t) x := by - refine' hg.mono fun x hx => _ - subst hx - simp only [mk_preimage_prod_right_fn_eq_if, indicator_apply, mem_preimage] - split_ifs - exacts [(map_apply hgm.of_uncurry_left ht).symm, measure_empty] - simp only [preimage_preimage] - rw [lintegral_congr_ae this, lintegral_indicator _ (hf.1 hs), set_lintegral_const, - hf.measure_preimage hs, mul_comm] + have : Measurable fun p : α × γ => (f p.1, g p.1 p.2) := (hf.1.comp measurable_fst).prod_mk hgm + /- if `μa = 0`, then the lemma is trivial, otherwise we can use `hg` + to deduce `SFinite μd`. -/ + rcases eq_or_ne μa 0 with (rfl | ha) + · rw [← hf.map_eq, zero_prod, Measure.map_zero, zero_prod] + exact ⟨this, by simp only [Measure.map_zero]⟩ + have sf : SFinite μd := by + rcases (ae_neBot.2 ha).nonempty_of_mem hg with ⟨x, hx : map (g x) μc = μd⟩ + rw [← hx] + infer_instance + -- Thus we can use the integral formula for the product measure, and compute things explicitly + refine ⟨this, ?_⟩ + ext s hs + rw [map_apply this hs, prod_apply (this hs), prod_apply hs, + ← hf.lintegral_comp (measurable_measure_prod_mk_left hs)] + apply lintegral_congr_ae + filter_upwards [hg] with a ha + rw [← ha, map_apply hgm.of_uncurry_left (measurable_prod_mk_left hs)] + rfl #align measure_theory.measure_preserving.skew_product MeasureTheory.MeasurePreserving.skew_product /-- If `f : α → β` sends the measure `μa` to `μb` and `g : γ → δ` sends the measure `μc` to `μd`, then `Prod.map f g` sends `μa.prod μc` to `μb.prod μd`. -/ -protected theorem prod [SigmaFinite μb] [SigmaFinite μd] {f : α → β} {g : γ → δ} +protected theorem prod [SFinite μa] [SFinite μc] {f : α → β} {g : γ → δ} (hf : MeasurePreserving f μa μb) (hg : MeasurePreserving g μc μd) : MeasurePreserving (Prod.map f g) (μa.prod μc) (μb.prod μd) := have : Measurable (uncurry fun _ : α => g) := hg.1.comp measurable_snd @@ -768,7 +838,7 @@ end MeasurePreserving namespace QuasiMeasurePreserving theorem prod_of_right {f : α × β → γ} {μ : Measure α} {ν : Measure β} {τ : Measure γ} - (hf : Measurable f) [SigmaFinite ν] + (hf : Measurable f) [SFinite ν] (h2f : ∀ᵐ x ∂μ, QuasiMeasurePreserving (fun y => f (x, y)) ν τ) : QuasiMeasurePreserving f (μ.prod ν) τ := by refine' ⟨hf, _⟩ @@ -779,7 +849,7 @@ theorem prod_of_right {f : α × β → γ} {μ : Measure α} {ν : Measure β} theorem prod_of_left {α β γ} [MeasurableSpace α] [MeasurableSpace β] [MeasurableSpace γ] {f : α × β → γ} {μ : Measure α} {ν : Measure β} {τ : Measure γ} (hf : Measurable f) - [SigmaFinite μ] [SigmaFinite ν] + [SFinite μ] [SFinite ν] (h2f : ∀ᵐ y ∂ν, QuasiMeasurePreserving (fun x => f (x, y)) μ τ) : QuasiMeasurePreserving f (μ.prod ν) τ := by rw [← prod_swap] @@ -796,18 +866,18 @@ open MeasureTheory.Measure section -theorem AEMeasurable.prod_swap [SigmaFinite μ] [SigmaFinite ν] {f : β × α → γ} +theorem AEMeasurable.prod_swap [SFinite μ] [SFinite ν] {f : β × α → γ} (hf : AEMeasurable f (ν.prod μ)) : AEMeasurable (fun z : α × β => f z.swap) (μ.prod ν) := by rw [← Measure.prod_swap] at hf exact hf.comp_measurable measurable_swap #align ae_measurable.prod_swap AEMeasurable.prod_swap -theorem AEMeasurable.fst [SigmaFinite ν] {f : α → γ} (hf : AEMeasurable f μ) : +theorem AEMeasurable.fst [SFinite ν] {f : α → γ} (hf : AEMeasurable f μ) : AEMeasurable (fun z : α × β => f z.1) (μ.prod ν) := hf.comp_quasiMeasurePreserving quasiMeasurePreserving_fst #align ae_measurable.fst AEMeasurable.fst -theorem AEMeasurable.snd [SigmaFinite ν] {f : β → γ} (hf : AEMeasurable f ν) : +theorem AEMeasurable.snd [SFinite ν] {f : β → γ} (hf : AEMeasurable f ν) : AEMeasurable (fun z : α × β => f z.2) (μ.prod ν) := hf.comp_quasiMeasurePreserving quasiMeasurePreserving_snd #align ae_measurable.snd AEMeasurable.snd @@ -819,9 +889,9 @@ namespace MeasureTheory /-! ### The Lebesgue integral on a product -/ -variable [SigmaFinite ν] +variable [SFinite ν] -theorem lintegral_prod_swap [SigmaFinite μ] (f : α × β → ℝ≥0∞) : +theorem lintegral_prod_swap [SFinite μ] (f : α × β → ℝ≥0∞) : ∫⁻ z, f z.swap ∂ν.prod μ = ∫⁻ z, f z ∂μ.prod ν := measurePreserving_swap.lintegral_comp_emb MeasurableEquiv.prodComm.measurableEmbedding f #align measure_theory.lintegral_prod_swap MeasureTheory.lintegral_prod_swap @@ -869,7 +939,7 @@ theorem lintegral_prod (f : α × β → ℝ≥0∞) (hf : AEMeasurable f (μ.pr /-- The symmetric version of Tonelli's Theorem: For `ℝ≥0∞`-valued almost everywhere measurable functions on `α × β`, the integral of `f` is equal to the iterated integral, in reverse order. -/ -theorem lintegral_prod_symm [SigmaFinite μ] (f : α × β → ℝ≥0∞) (hf : AEMeasurable f (μ.prod ν)) : +theorem lintegral_prod_symm [SFinite μ] (f : α × β → ℝ≥0∞) (hf : AEMeasurable f (μ.prod ν)) : ∫⁻ z, f z ∂μ.prod ν = ∫⁻ y, ∫⁻ x, f (x, y) ∂μ ∂ν := by simp_rw [← lintegral_prod_swap f] exact lintegral_prod _ hf.prod_swap @@ -877,7 +947,7 @@ theorem lintegral_prod_symm [SigmaFinite μ] (f : α × β → ℝ≥0∞) (hf : /-- The symmetric version of Tonelli's Theorem: For `ℝ≥0∞`-valued measurable functions on `α × β`, the integral of `f` is equal to the iterated integral, in reverse order. -/ -theorem lintegral_prod_symm' [SigmaFinite μ] (f : α × β → ℝ≥0∞) (hf : Measurable f) : +theorem lintegral_prod_symm' [SFinite μ] (f : α × β → ℝ≥0∞) (hf : Measurable f) : ∫⁻ z, f z ∂μ.prod ν = ∫⁻ y, ∫⁻ x, f (x, y) ∂μ ∂ν := lintegral_prod_symm f hf.aemeasurable #align measure_theory.lintegral_prod_symm' MeasureTheory.lintegral_prod_symm' @@ -891,14 +961,14 @@ theorem lintegral_lintegral ⦃f : α → β → ℝ≥0∞⦄ (hf : AEMeasurabl /-- The reversed version of **Tonelli's Theorem** (symmetric version). In this version `f` is in curried form, which makes it easier for the elaborator to figure out `f` automatically. -/ -theorem lintegral_lintegral_symm [SigmaFinite μ] ⦃f : α → β → ℝ≥0∞⦄ +theorem lintegral_lintegral_symm [SFinite μ] ⦃f : α → β → ℝ≥0∞⦄ (hf : AEMeasurable (uncurry f) (μ.prod ν)) : ∫⁻ x, ∫⁻ y, f x y ∂ν ∂μ = ∫⁻ z, f z.2 z.1 ∂ν.prod μ := (lintegral_prod_symm _ hf.prod_swap).symm #align measure_theory.lintegral_lintegral_symm MeasureTheory.lintegral_lintegral_symm /-- Change the order of Lebesgue integration. -/ -theorem lintegral_lintegral_swap [SigmaFinite μ] ⦃f : α → β → ℝ≥0∞⦄ +theorem lintegral_lintegral_swap [SFinite μ] ⦃f : α → β → ℝ≥0∞⦄ (hf : AEMeasurable (uncurry f) (μ.prod ν)) : ∫⁻ x, ∫⁻ y, f x y ∂ν ∂μ = ∫⁻ y, ∫⁻ x, f x y ∂μ ∂ν := (lintegral_lintegral hf).trans (lintegral_prod_symm _ hf) diff --git a/Mathlib/MeasureTheory/Covering/BesicovitchVectorSpace.lean b/Mathlib/MeasureTheory/Covering/BesicovitchVectorSpace.lean index a319bcfcafd12..4941e0635489a 100644 --- a/Mathlib/MeasureTheory/Covering/BesicovitchVectorSpace.lean +++ b/Mathlib/MeasureTheory/Covering/BesicovitchVectorSpace.lean @@ -210,7 +210,7 @@ theorem exists_goodδ : subsequence, to obtain a `1`-separated set in the ball of radius `2` with cardinality `N = multiplicity E + 1`. To formalize this, we work with functions `Fin N → E`. -/ - by_contra' h + by_contra! h set N := multiplicity E + 1 with hN have : ∀ δ : ℝ, 0 < δ → ∃ f : Fin N → E, (∀ i : Fin N, ‖f i‖ ≤ 2) ∧ diff --git a/Mathlib/MeasureTheory/Decomposition/SignedHahn.lean b/Mathlib/MeasureTheory/Decomposition/SignedHahn.lean index 48cdbbafd0d46..2548a786a2ba2 100644 --- a/Mathlib/MeasureTheory/Decomposition/SignedHahn.lean +++ b/Mathlib/MeasureTheory/Decomposition/SignedHahn.lean @@ -342,7 +342,7 @@ theorem zero_mem_measureOfNegatives : (0 : ℝ) ∈ s.measureOfNegatives := theorem bddBelow_measureOfNegatives : BddBelow s.measureOfNegatives := by simp_rw [BddBelow, Set.Nonempty, mem_lowerBounds] - by_contra' h + by_contra! h have h' : ∀ n : ℕ, ∃ y : ℝ, y ∈ s.measureOfNegatives ∧ y < -n := fun n => h (-n) choose f hf using h' have hf' : ∀ n : ℕ, ∃ B, MeasurableSet B ∧ s ≤[B] 0 ∧ s B < -n := by @@ -395,7 +395,7 @@ theorem exists_compl_positive_negative (s : SignedMeasure α) : refine' ⟨Aᶜ, hA₁.compl, _, (compl_compl A).symm ▸ hA₂⟩ rw [restrict_le_restrict_iff _ _ hA₁.compl] intro C _ hC₁ - by_contra' hC₂ + by_contra! hC₂ rcases exists_subset_restrict_nonpos hC₂ with ⟨D, hD₁, hD, hD₂, hD₃⟩ have : s (A ∪ D) < sInf s.measureOfNegatives := by rw [← hA₃, diff --git a/Mathlib/MeasureTheory/Function/LpSpace.lean b/Mathlib/MeasureTheory/Function/LpSpace.lean index 8130d549ec91f..f268c7ffb6fdb 100644 --- a/Mathlib/MeasureTheory/Function/LpSpace.lean +++ b/Mathlib/MeasureTheory/Function/LpSpace.lean @@ -574,7 +574,7 @@ theorem snormEssSup_indicator_const_le (s : Set α) (c : G) : theorem snormEssSup_indicator_const_eq (s : Set α) (c : G) (hμs : μ s ≠ 0) : snormEssSup (s.indicator fun _ : α => c) μ = ‖c‖₊ := by refine' le_antisymm (snormEssSup_indicator_const_le s c) _ - by_contra' h + by_contra! h have h' := ae_iff.mp (ae_lt_of_essSup_lt h) push_neg at h' refine' hμs (measure_mono_null (fun x hx_mem => _) h') diff --git a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean index 58b8060a254cf..7dcb794bcdbd9 100644 --- a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean +++ b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean @@ -546,14 +546,14 @@ open Filter protected theorem sup [Sup β] [ContinuousSup β] (hf : StronglyMeasurable f) (hg : StronglyMeasurable g) : StronglyMeasurable (f ⊔ g) := ⟨fun n => hf.approx n ⊔ hg.approx n, fun x => - (hf.tendsto_approx x).sup_right_nhds (hg.tendsto_approx x)⟩ + (hf.tendsto_approx x).sup_nhds (hg.tendsto_approx x)⟩ #align measure_theory.strongly_measurable.sup MeasureTheory.StronglyMeasurable.sup @[aesop safe 20 (rule_sets [Measurable])] protected theorem inf [Inf β] [ContinuousInf β] (hf : StronglyMeasurable f) (hg : StronglyMeasurable g) : StronglyMeasurable (f ⊓ g) := ⟨fun n => hf.approx n ⊓ hg.approx n, fun x => - (hf.tendsto_approx x).inf_right_nhds (hg.tendsto_approx x)⟩ + (hf.tendsto_approx x).inf_nhds (hg.tendsto_approx x)⟩ #align measure_theory.strongly_measurable.inf MeasureTheory.StronglyMeasurable.inf end Order @@ -1120,7 +1120,7 @@ protected theorem sup [SemilatticeSup β] [ContinuousSup β] (hf : FinStronglyMe (hg : FinStronglyMeasurable g μ) : FinStronglyMeasurable (f ⊔ g) μ := by refine' ⟨fun n => hf.approx n ⊔ hg.approx n, fun n => _, fun x => - (hf.tendsto_approx x).sup_right_nhds (hg.tendsto_approx x)⟩ + (hf.tendsto_approx x).sup_nhds (hg.tendsto_approx x)⟩ refine' (measure_mono (support_sup _ _)).trans_lt _ exact measure_union_lt_top_iff.mpr ⟨hf.fin_support_approx n, hg.fin_support_approx n⟩ #align measure_theory.fin_strongly_measurable.sup MeasureTheory.FinStronglyMeasurable.sup @@ -1130,7 +1130,7 @@ protected theorem inf [SemilatticeInf β] [ContinuousInf β] (hf : FinStronglyMe (hg : FinStronglyMeasurable g μ) : FinStronglyMeasurable (f ⊓ g) μ := by refine' ⟨fun n => hf.approx n ⊓ hg.approx n, fun n => _, fun x => - (hf.tendsto_approx x).inf_right_nhds (hg.tendsto_approx x)⟩ + (hf.tendsto_approx x).inf_nhds (hg.tendsto_approx x)⟩ refine' (measure_mono (support_inf _ _)).trans_lt _ exact measure_union_lt_top_iff.mpr ⟨hf.fin_support_approx n, hg.fin_support_approx n⟩ #align measure_theory.fin_strongly_measurable.inf MeasureTheory.FinStronglyMeasurable.inf diff --git a/Mathlib/MeasureTheory/Group/Action.lean b/Mathlib/MeasureTheory/Group/Action.lean index f7a5d5cee19a7..ed881629fa6ba 100644 --- a/Mathlib/MeasureTheory/Group/Action.lean +++ b/Mathlib/MeasureTheory/Group/Action.lean @@ -121,7 +121,7 @@ theorem smulInvariantMeasure_map [SMul M α] [SMul M β] _ = μ (f ⁻¹' ((m • ·) ⁻¹' S)) := map_apply hf <| hS.preimage (measurable_const_smul _) _ = μ ((m • f ·) ⁻¹' S) := by rw [preimage_preimage] _ = μ ((f <| m • ·) ⁻¹' S) := by simp_rw [hsmul] - _ = μ ((m • ·) ⁻¹' (f ⁻¹' S)) := by rw [←preimage_preimage] + _ = μ ((m • ·) ⁻¹' (f ⁻¹' S)) := by rw [← preimage_preimage] _ = μ (f ⁻¹' S) := by rw [SMulInvariantMeasure.measure_preimage_smul m (hS.preimage hf)] _ = map f μ S := (map_apply hf hS).symm diff --git a/Mathlib/MeasureTheory/Group/AddCircle.lean b/Mathlib/MeasureTheory/Group/AddCircle.lean index b821dfa8fe618..861b5ac182994 100644 --- a/Mathlib/MeasureTheory/Group/AddCircle.lean +++ b/Mathlib/MeasureTheory/Group/AddCircle.lean @@ -83,7 +83,7 @@ theorem isAddFundamentalDomain_of_ae_ball (I : Set <| AddCircle T) (u x : AddCir haveI : Fintype G := @Fintype.ofFinite _ hu.finite_zmultiples.to_subtype have hG_card : (Finset.univ : Finset G).card = n := by show _ = addOrderOf u - rw [←Nat.card_zmultiples, Nat.card_eq_fintype_card]; rfl + rw [← Nat.card_zmultiples, Nat.card_eq_fintype_card]; rfl simp_rw [measure_vadd] rw [AddCircle.measure_univ, tsum_fintype, Finset.sum_const, measure_congr hI, volume_closedBall, ← ENNReal.ofReal_nsmul, mul_div, mul_div_mul_comm, @@ -101,7 +101,7 @@ theorem volume_of_add_preimage_eq (s I : Set <| AddCircle T) (u x : AddCircle T) have hsG : ∀ g : G, (g +ᵥ s : Set <| AddCircle T) =ᵐ[volume] s := by rintro ⟨y, hy⟩; exact (vadd_ae_eq_self_of_mem_zmultiples hs hy : _) rw [(isAddFundamentalDomain_of_ae_ball I u x hu hI).measure_eq_card_smul_of_vadd_ae_eq_self s hsG, - ←Nat.card_zmultiples u, Nat.card_eq_fintype_card] + ← Nat.card_zmultiples u, Nat.card_eq_fintype_card] #align add_circle.volume_of_add_preimage_eq AddCircle.volume_of_add_preimage_eq end AddCircle diff --git a/Mathlib/MeasureTheory/Group/Measure.lean b/Mathlib/MeasureTheory/Group/Measure.lean index 4f79d9acdb66d..d12de64415911 100644 --- a/Mathlib/MeasureTheory/Group/Measure.lean +++ b/Mathlib/MeasureTheory/Group/Measure.lean @@ -202,30 +202,26 @@ theorem forall_measure_preimage_mul_right_iff (μ : Measure G) : #align measure_theory.forall_measure_preimage_add_right_iff MeasureTheory.forall_measure_preimage_add_right_iff @[to_additive] -instance Measure.prod.instIsMulLeftInvariant [IsMulLeftInvariant μ] [SigmaFinite μ] {H : Type*} +instance Measure.prod.instIsMulLeftInvariant [IsMulLeftInvariant μ] [SFinite μ] {H : Type*} [Mul H] {mH : MeasurableSpace H} {ν : Measure H} [MeasurableMul H] [IsMulLeftInvariant ν] - [SigmaFinite ν] : IsMulLeftInvariant (μ.prod ν) := by + [SFinite ν] : IsMulLeftInvariant (μ.prod ν) := by constructor rintro ⟨g, h⟩ change map (Prod.map (g * ·) (h * ·)) (μ.prod ν) = μ.prod ν rw [← map_prod_map _ _ (measurable_const_mul g) (measurable_const_mul h), map_mul_left_eq_self μ g, map_mul_left_eq_self ν h] - · rw [map_mul_left_eq_self μ g]; infer_instance - · rw [map_mul_left_eq_self ν h]; infer_instance #align measure_theory.measure.prod.measure.is_mul_left_invariant MeasureTheory.Measure.prod.instIsMulLeftInvariant #align measure_theory.measure.prod.measure.is_add_left_invariant MeasureTheory.Measure.prod.instIsAddLeftInvariant @[to_additive] -instance Measure.prod.instIsMulRightInvariant [IsMulRightInvariant μ] [SigmaFinite μ] {H : Type*} +instance Measure.prod.instIsMulRightInvariant [IsMulRightInvariant μ] [SFinite μ] {H : Type*} [Mul H] {mH : MeasurableSpace H} {ν : Measure H} [MeasurableMul H] [IsMulRightInvariant ν] - [SigmaFinite ν] : IsMulRightInvariant (μ.prod ν) := by + [SFinite ν] : IsMulRightInvariant (μ.prod ν) := by constructor rintro ⟨g, h⟩ change map (Prod.map (· * g) (· * h)) (μ.prod ν) = μ.prod ν rw [← map_prod_map _ _ (measurable_mul_const g) (measurable_mul_const h), map_mul_right_eq_self μ g, map_mul_right_eq_self ν h] - · rw [map_mul_right_eq_self μ g]; infer_instance - · rw [map_mul_right_eq_self ν h]; infer_instance #align measure_theory.measure.prod.measure.is_mul_right_invariant MeasureTheory.Measure.prod.instIsMulRightInvariant #align measure_theory.measure.prod.measure.is_add_right_invariant MeasureTheory.Measure.prod.instIsMulRightInvariant @@ -435,6 +431,10 @@ theorem measurePreserving_inv (μ : Measure G) [IsInvInvariant μ] : MeasurePres #align measure_theory.measure.measure_preserving_inv MeasureTheory.Measure.measurePreserving_inv #align measure_theory.measure.measure_preserving_neg MeasureTheory.Measure.measurePreserving_neg +@[to_additive] +instance inv.instSFinite (μ : Measure G) [SFinite μ] : SFinite μ.inv := by + rw [Measure.inv]; infer_instance + end Inv section InvolutiveInv @@ -954,7 +954,7 @@ instance (priority := 100) IsHaarMeasure.sigmaFinite [SigmaCompactSpace G] : Sig @[to_additive] instance prod.instIsHaarMeasure {G : Type*} [Group G] [TopologicalSpace G] {_ : MeasurableSpace G} {H : Type*} [Group H] [TopologicalSpace H] {_ : MeasurableSpace H} (μ : Measure G) - (ν : Measure H) [IsHaarMeasure μ] [IsHaarMeasure ν] [SigmaFinite μ] [SigmaFinite ν] + (ν : Measure H) [IsHaarMeasure μ] [IsHaarMeasure ν] [SFinite μ] [SFinite ν] [MeasurableMul G] [MeasurableMul H] : IsHaarMeasure (μ.prod ν) where #align measure_theory.measure.prod.is_haar_measure MeasureTheory.Measure.prod.instIsHaarMeasure #align measure_theory.measure.prod.is_add_haar_measure MeasureTheory.Measure.prod.instIsAddHaarMeasure diff --git a/Mathlib/MeasureTheory/Integral/Average.lean b/Mathlib/MeasureTheory/Integral/Average.lean index d27e9b6cea67f..707385eba14de 100644 --- a/Mathlib/MeasureTheory/Integral/Average.lean +++ b/Mathlib/MeasureTheory/Integral/Average.lean @@ -149,14 +149,14 @@ theorem laverage_add_measure : · rw [not_isFiniteMeasure_iff] at hν simp [laverage_eq, hν] haveI := hμ; haveI := hν - simp only [←ENNReal.mul_div_right_comm, measure_mul_laverage, ←ENNReal.add_div, - ←lintegral_add_measure, ←Measure.add_apply, ←laverage_eq] + simp only [← ENNReal.mul_div_right_comm, measure_mul_laverage, ← ENNReal.add_div, + ← lintegral_add_measure, ← Measure.add_apply, ← laverage_eq] #align measure_theory.laverage_add_measure MeasureTheory.laverage_add_measure theorem measure_mul_setLaverage (f : α → ℝ≥0∞) (h : μ s ≠ ∞) : μ s * ⨍⁻ x in s, f x ∂μ = ∫⁻ x in s, f x ∂μ := by have := Fact.mk h.lt_top - rw [←measure_mul_laverage, restrict_apply_univ] + rw [← measure_mul_laverage, restrict_apply_univ] #align measure_theory.measure_mul_set_laverage MeasureTheory.measure_mul_setLaverage theorem laverage_union (hd : AEDisjoint μ s t) (ht : NullMeasurableSet t μ) : @@ -171,7 +171,7 @@ theorem laverage_union_mem_openSegment (hd : AEDisjoint μ s t) (ht : NullMeasur refine' ⟨μ s / (μ s + μ t), μ t / (μ s + μ t), ENNReal.div_pos hs₀ <| add_ne_top.2 ⟨hsμ, htμ⟩, ENNReal.div_pos ht₀ <| add_ne_top.2 ⟨hsμ, htμ⟩, _, (laverage_union hd ht).symm⟩ - rw [←ENNReal.add_div, + rw [← ENNReal.add_div, ENNReal.div_self (add_eq_zero.not.2 fun h => hs₀ h.1) (add_ne_top.2 ⟨hsμ, htμ⟩)] #align measure_theory.laverage_union_mem_open_segment MeasureTheory.laverage_union_mem_openSegment @@ -179,12 +179,12 @@ theorem laverage_union_mem_segment (hd : AEDisjoint μ s t) (ht : NullMeasurable (hsμ : μ s ≠ ∞) (htμ : μ t ≠ ∞) : ⨍⁻ x in s ∪ t, f x ∂μ ∈ [⨍⁻ x in s, f x ∂μ -[ℝ≥0∞] ⨍⁻ x in t, f x ∂μ] := by by_cases hs₀ : μ s = 0 - · rw [←ae_eq_empty] at hs₀ + · rw [← ae_eq_empty] at hs₀ rw [restrict_congr_set (hs₀.union EventuallyEq.rfl), empty_union] exact right_mem_segment _ _ _ · refine' ⟨μ s / (μ s + μ t), μ t / (μ s + μ t), zero_le _, zero_le _, _, (laverage_union hd ht).symm⟩ - rw [←ENNReal.add_div, + rw [← ENNReal.add_div, ENNReal.div_self (add_eq_zero.not.2 fun h => hs₀ h.1) (add_ne_top.2 ⟨hsμ, htμ⟩)] #align measure_theory.laverage_union_mem_segment MeasureTheory.laverage_union_mem_segment @@ -324,8 +324,8 @@ theorem average_add_measure [IsFiniteMeasure μ] {ν : Measure α} [IsFiniteMeas ⨍ x, f x ∂(μ + ν) = ((μ univ).toReal / ((μ univ).toReal + (ν univ).toReal)) • ⨍ x, f x ∂μ + ((ν univ).toReal / ((μ univ).toReal + (ν univ).toReal)) • ⨍ x, f x ∂ν := by - simp only [div_eq_inv_mul, mul_smul, measure_smul_average, ←smul_add, - ←integral_add_measure hμ hν, ←ENNReal.toReal_add (measure_ne_top μ _) (measure_ne_top ν _)] + simp only [div_eq_inv_mul, mul_smul, measure_smul_average, ← smul_add, + ← integral_add_measure hμ hν, ← ENNReal.toReal_add (measure_ne_top μ _) (measure_ne_top ν _)] rw [average_eq, Measure.add_apply] #align measure_theory.average_add_measure MeasureTheory.average_add_measure @@ -337,7 +337,7 @@ theorem average_pair {f : α → E} {g : α → F} (hfi : Integrable f μ) (hgi theorem measure_smul_setAverage (f : α → E) {s : Set α} (h : μ s ≠ ∞) : (μ s).toReal • ⨍ x in s, f x ∂μ = ∫ x in s, f x ∂μ := by haveI := Fact.mk h.lt_top - rw [←measure_smul_average, restrict_apply_univ] + rw [← measure_smul_average, restrict_apply_univ] #align measure_theory.measure_smul_set_average MeasureTheory.measure_smul_setAverage theorem average_union {f : α → E} {s t : Set α} (hd : AEDisjoint μ s t) (ht : NullMeasurableSet t μ) @@ -364,7 +364,7 @@ theorem average_union_mem_segment {f : α → E} {s t : Set α} (hd : AEDisjoint (hft : IntegrableOn f t μ) : ⨍ x in s ∪ t, f x ∂μ ∈ [⨍ x in s, f x ∂μ -[ℝ] ⨍ x in t, f x ∂μ] := by by_cases hse : μ s = 0 - · rw [←ae_eq_empty] at hse + · rw [← ae_eq_empty] at hse rw [restrict_congr_set (hse.union EventuallyEq.rfl), empty_union] exact right_mem_segment _ _ _ · refine' @@ -613,15 +613,15 @@ theorem measure_le_setLaverage_pos (hμ : μ s ≠ 0) (hμ₁ : μ s ≠ ∞) obtain h | h := eq_or_ne (∫⁻ a in s, f a ∂μ) ∞ · simpa [mul_top, hμ₁, laverage, h, top_div_of_ne_top hμ₁, pos_iff_ne_zero] using hμ have := measure_le_setAverage_pos hμ hμ₁ (integrable_toReal_of_lintegral_ne_top hf h) - rw [←setOf_inter_eq_sep, ←Measure.restrict_apply₀ + rw [← setOf_inter_eq_sep, ← Measure.restrict_apply₀ (hf.aestronglyMeasurable.nullMeasurableSet_le aestronglyMeasurable_const)] - rw [←setOf_inter_eq_sep, ←Measure.restrict_apply₀ + rw [← setOf_inter_eq_sep, ← Measure.restrict_apply₀ (hf.ennreal_toReal.aestronglyMeasurable.nullMeasurableSet_le aestronglyMeasurable_const), - ←measure_diff_null (measure_eq_top_of_lintegral_ne_top hf h)] at this + ← measure_diff_null (measure_eq_top_of_lintegral_ne_top hf h)] at this refine' this.trans_le (measure_mono _) rintro x ⟨hfx, hx⟩ dsimp at hfx - rwa [←toReal_laverage hf, toReal_le_toReal hx (setLaverage_lt_top h).ne] at hfx + rwa [← toReal_laverage hf, toReal_le_toReal hx (setLaverage_lt_top h).ne] at hfx · simp_rw [ae_iff, not_ne_iff] exact measure_eq_top_of_lintegral_ne_top hf h #align measure_theory.measure_le_set_laverage_pos MeasureTheory.measure_le_setLaverage_pos @@ -637,13 +637,13 @@ theorem measure_setLaverage_le_pos (hμ : μ s ≠ 0) (hs : NullMeasurableSet s rw [hfg] at hint have := measure_setAverage_le_pos hμ hμ₁ (integrable_toReal_of_lintegral_ne_top hg.aemeasurable hint) - simp_rw [←setOf_inter_eq_sep, ←Measure.restrict_apply₀' hs, hfg'] - rw [←setOf_inter_eq_sep, ←Measure.restrict_apply₀' hs, ← + simp_rw [← setOf_inter_eq_sep, ← Measure.restrict_apply₀' hs, hfg'] + rw [← setOf_inter_eq_sep, ← Measure.restrict_apply₀' hs, ← measure_diff_null (measure_eq_top_of_lintegral_ne_top hg.aemeasurable hint)] at this refine' this.trans_le (measure_mono _) rintro x ⟨hfx, hx⟩ dsimp at hfx - rw [←toReal_laverage hg.aemeasurable, toReal_le_toReal (setLaverage_lt_top hint).ne hx] at hfx + rw [← toReal_laverage hg.aemeasurable, toReal_le_toReal (setLaverage_lt_top hint).ne hx] at hfx exact hfx.trans (hgf _) · simp_rw [ae_iff, not_ne_iff] exact measure_eq_top_of_lintegral_ne_top hg.aemeasurable hint @@ -682,7 +682,7 @@ avoiding a null set. -/ theorem exists_not_mem_null_laverage_le (hμ : μ ≠ 0) (hint : ∫⁻ a : α, f a ∂μ ≠ ∞) (hN : μ N = 0) : ∃ x, x ∉ N ∧ ⨍⁻ a, f a ∂μ ≤ f x := by have := measure_laverage_le_pos hμ hint - rw [←measure_diff_null hN] at this + rw [← measure_diff_null hN] at this obtain ⟨x, hx, hxN⟩ := nonempty_of_measure_ne_zero this.ne' exact ⟨x, hxN, hx⟩ #align measure_theory.exists_not_mem_null_laverage_le MeasureTheory.exists_not_mem_null_laverage_le @@ -709,7 +709,7 @@ avoiding a null set. -/ theorem exists_not_mem_null_le_laverage (hμ : μ ≠ 0) (hf : AEMeasurable f μ) (hN : μ N = 0) : ∃ x, x ∉ N ∧ f x ≤ ⨍⁻ a, f a ∂μ := by have := measure_le_laverage_pos hμ hf - rw [←measure_diff_null hN] at this + rw [← measure_diff_null hN] at this obtain ⟨x, hx, hxN⟩ := nonempty_of_measure_ne_zero this.ne' exact ⟨x, hxN, hx⟩ #align measure_theory.exists_not_mem_null_le_laverage MeasureTheory.exists_not_mem_null_le_laverage diff --git a/Mathlib/MeasureTheory/Integral/FundThmCalculus.lean b/Mathlib/MeasureTheory/Integral/FundThmCalculus.lean index 7b36a4391a5b5..87797dbeb4adb 100644 --- a/Mathlib/MeasureTheory/Integral/FundThmCalculus.lean +++ b/Mathlib/MeasureTheory/Integral/FundThmCalculus.lean @@ -1261,7 +1261,7 @@ theorem integrableOn_deriv_right_of_nonneg (hcont : ContinuousOn g (Icc a b)) exact (hderiv x hx).derivWithin (uniqueDiffWithinAt_Ioi _) suffices H : (∫⁻ x in Ioo a b, ‖g' x‖₊) ≤ ENNReal.ofReal (g b - g a) from ⟨meas_g'.aestronglyMeasurable, H.trans_lt ENNReal.ofReal_lt_top⟩ - by_contra' H + by_contra! H obtain ⟨f, fle, fint, hf⟩ : ∃ f : SimpleFunc ℝ ℝ≥0, (∀ x, f x ≤ ‖g' x‖₊) ∧ diff --git a/Mathlib/MeasureTheory/Integral/Lebesgue.lean b/Mathlib/MeasureTheory/Integral/Lebesgue.lean index ac5432f5cf20e..bdd56bf4e6943 100644 --- a/Mathlib/MeasureTheory/Integral/Lebesgue.lean +++ b/Mathlib/MeasureTheory/Integral/Lebesgue.lean @@ -95,7 +95,7 @@ theorem lintegral_mono' {m : MeasurableSpace α} ⦃μ ν : Measure α⦄ (hμν -- workaround for the known eta-reduction issue with `@[gcongr]` @[gcongr] theorem lintegral_mono_fn' ⦃f g : α → ℝ≥0∞⦄ (hfg : ∀ x, f x ≤ g x) (h2 : μ ≤ ν) : lintegral μ f ≤ lintegral ν g := -lintegral_mono' h2 hfg + lintegral_mono' h2 hfg theorem lintegral_mono ⦃f g : α → ℝ≥0∞⦄ (hfg : f ≤ g) : ∫⁻ a, f a ∂μ ≤ ∫⁻ a, g a ∂μ := lintegral_mono' (le_refl μ) hfg @@ -104,7 +104,7 @@ theorem lintegral_mono ⦃f g : α → ℝ≥0∞⦄ (hfg : f ≤ g) : ∫⁻ a, -- workaround for the known eta-reduction issue with `@[gcongr]` @[gcongr] theorem lintegral_mono_fn ⦃f g : α → ℝ≥0∞⦄ (hfg : ∀ x, f x ≤ g x) : lintegral μ f ≤ lintegral μ g := -lintegral_mono hfg + lintegral_mono hfg theorem lintegral_mono_nnreal {f g : α → ℝ≥0} (h : f ≤ g) : ∫⁻ a, f a ∂μ ≤ ∫⁻ a, g a ∂μ := lintegral_mono fun a => ENNReal.coe_le_coe.2 (h a) @@ -884,7 +884,7 @@ theorem lintegral_eq_top_of_measure_eq_top_ne_zero {f : α → ℝ≥0∞} (hf : theorem setLintegral_eq_top_of_measure_eq_top_ne_zero (hf : AEMeasurable f (μ.restrict s)) (hμf : μ ({x ∈ s | f x = ∞}) ≠ 0) : ∫⁻ x in s, f x ∂μ = ∞ := lintegral_eq_top_of_measure_eq_top_ne_zero hf $ - mt (eq_bot_mono $ by rw [←setOf_inter_eq_sep]; exact Measure.le_restrict_apply _ _) hμf + mt (eq_bot_mono $ by rw [← setOf_inter_eq_sep]; exact Measure.le_restrict_apply _ _) hμf #align measure_theory.set_lintegral_eq_top_of_measure_eq_top_ne_zero MeasureTheory.setLintegral_eq_top_of_measure_eq_top_ne_zero theorem measure_eq_top_of_lintegral_ne_top (hf : AEMeasurable f μ) (hμf : ∫⁻ x, f x ∂μ ≠ ∞) : diff --git a/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean b/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean index ba5a37cd8ba25..feec30b6246d1 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean @@ -1373,6 +1373,10 @@ attribute [simps! apply toEquiv] trans refl @[simp] theorem symm_symm (e : α ≃ᵐ β) : e.symm.symm = e := rfl +theorem symm_bijective : + Function.Bijective (MeasurableEquiv.symm : (α ≃ᵐ β) → β ≃ᵐ α) := + Function.bijective_iff_has_inverse.mpr ⟨_, symm_symm, symm_symm⟩ + @[simp] theorem symm_refl (α : Type*) [MeasurableSpace α] : (refl α).symm = refl α := rfl @@ -1861,7 +1865,7 @@ theorem MeasurableSpace.comap_compl {m' : MeasurableSpace β} [BooleanAlgebra β (h : Measurable (compl : β → β)) (f : α → β) : MeasurableSpace.comap (fun a => (f a)ᶜ) inferInstance = MeasurableSpace.comap f inferInstance := by - rw [←Function.comp_def, ←MeasurableSpace.comap_comp] + rw [← Function.comp_def, ← MeasurableSpace.comap_comp] congr exact (MeasurableEquiv.ofInvolutive _ compl_involutive h).measurableEmbedding.comap_eq #align measurable_space.comap_compl MeasurableSpace.comap_compl diff --git a/Mathlib/MeasureTheory/Measure/AEMeasurable.lean b/Mathlib/MeasureTheory/Measure/AEMeasurable.lean index 6ad05c2e2f51a..8282d50ca493e 100644 --- a/Mathlib/MeasureTheory/Measure/AEMeasurable.lean +++ b/Mathlib/MeasureTheory/Measure/AEMeasurable.lean @@ -15,7 +15,6 @@ function. This property, called `AEMeasurable f μ`, is defined in the file `Mea We discuss several of its properties that are analogous to properties of measurable functions. -/ - open MeasureTheory MeasureTheory.Measure Filter Set Function Classical ENNReal variable {ι α β γ δ R : Type*} {m0 : MeasurableSpace α} [MeasurableSpace β] [MeasurableSpace γ] @@ -447,3 +446,24 @@ lemma MeasureTheory.NullMeasurable.aemeasurable_of_aerange {f : α → β} {t : lift f' to α → t using hf't replace hf'm : NullMeasurable f' μ := hf'm.measurable'.subtype_mk exact (measurable_subtype_coe.comp_aemeasurable hf'm.aemeasurable).congr hff'.symm + +namespace MeasureTheory +namespace Measure + +lemma map_sum {ι : Type*} {m : ι → Measure α} {f : α → β} (hf : AEMeasurable f (Measure.sum m)) : + Measure.map f (Measure.sum m) = Measure.sum (fun i ↦ Measure.map f (m i)) := by + ext s hs + rw [map_apply_of_aemeasurable hf hs, sum_apply₀ _ (hf.nullMeasurable hs), sum_apply _ hs] + have M i : AEMeasurable f (m i) := hf.mono_measure (le_sum m i) + simp_rw [map_apply_of_aemeasurable (M _) hs] + +instance (μ : Measure α) (f : α → β) [SFinite μ] : SFinite (μ.map f) := by + by_cases H : AEMeasurable f μ + · rw [← sum_sFiniteSeq μ] at H ⊢ + rw [map_sum H] + infer_instance + · rw [map_of_not_aemeasurable H] + infer_instance + +end Measure +end MeasureTheory diff --git a/Mathlib/MeasureTheory/Measure/FiniteMeasure.lean b/Mathlib/MeasureTheory/Measure/FiniteMeasure.lean index 119e882c17f42..42e81dbc00b8a 100644 --- a/Mathlib/MeasureTheory/Measure/FiniteMeasure.lean +++ b/Mathlib/MeasureTheory/Measure/FiniteMeasure.lean @@ -5,6 +5,7 @@ Authors: Kalle Kytölä -/ import Mathlib.Topology.Algebra.Module.WeakDual import Mathlib.MeasureTheory.Integral.BoundedContinuousFunction +import Mathlib.MeasureTheory.Measure.HasOuterApproxClosed #align_import measure_theory.measure.finite_measure from "leanprover-community/mathlib"@"f2ce6086713c78a7f880485f7917ea547a215982" @@ -45,6 +46,9 @@ The main definitions are `MeasureTheory.FiniteMeasure.tendsto_iff_forall_lintegral_tendsto`. * `MeasureTheory.FiniteMeasure.continuous_map`: For a continuous function `f : Ω → Ω'`, the push-forward of finite measures `f* : FiniteMeasure Ω → FiniteMeasure Ω'` is continuous. + * `MeasureTheory.FiniteMeasure.t2Space`: The topology of weak convergence of finite Borel measures + is Hausdorff on spaces where indicators of closed sets have continuous decreasing approximating + sequences (in particular on any pseudo-metrizable spaces). ## Implementation notes @@ -315,6 +319,15 @@ theorem restrict_nonzero_iff (μ : FiniteMeasure Ω) (A : Set Ω) : μ.restrict variable [TopologicalSpace Ω] +/-- Two finite Borel measures are equal if the integrals of all bounded continuous functions with +respect to both agree. -/ +theorem ext_of_forall_lintegral_eq [HasOuterApproxClosed Ω] [BorelSpace Ω] + {μ ν : FiniteMeasure Ω} (h : ∀ (f : Ω →ᵇ ℝ≥0), ∫⁻ x, f x ∂μ = ∫⁻ x, f x ∂ν) : + μ = ν := by + apply Subtype.ext + change (μ : Measure Ω) = (ν : Measure Ω) + exact ext_of_forall_lintegral_eq_of_IsFiniteMeasure h + /-- The pairing of a finite (Borel) measure `μ` with a nonnegative bounded continuous function is obtained by (Lebesgue) integrating the (test) function against the measure. This is `MeasureTheory.FiniteMeasure.testAgainstNN`. -/ @@ -368,6 +381,8 @@ theorem smul_testAgainstNN_apply (c : ℝ≥0) (μ : FiniteMeasure Ω) (f : Ω lintegral_smul_measure] #align measure_theory.finite_measure.smul_test_against_nn_apply MeasureTheory.FiniteMeasure.smul_testAgainstNN_apply +section weak_convergence + variable [OpensMeasurableSpace Ω] theorem testAgainstNN_add (μ : FiniteMeasure Ω) (f₁ f₂ : Ω →ᵇ ℝ≥0) : @@ -540,6 +555,42 @@ theorem tendsto_iff_forall_lintegral_tendsto {γ : Type*} {F : Filter γ} {μs : ENNReal.toNNReal_coe] #align measure_theory.finite_measure.tendsto_iff_forall_lintegral_tendsto MeasureTheory.FiniteMeasure.tendsto_iff_forall_lintegral_tendsto +end weak_convergence -- section + +section Hausdorff + +variable [HasOuterApproxClosed Ω] [BorelSpace Ω] + +open Function + +/-- The mapping `toWeakDualBCNN` from finite Borel measures to the weak dual of `Ω →ᵇ ℝ≥0` is +injective, if in the underlying space `Ω`, indicator functions of closed sets have decreasing +approximations by sequences of continuous functions (in particular if `Ω` is pseudometrizable). -/ +lemma injective_toWeakDualBCNN : + Injective (toWeakDualBCNN : FiniteMeasure Ω → WeakDual ℝ≥0 (Ω →ᵇ ℝ≥0)) := by + intro μ ν hμν + apply ext_of_forall_lintegral_eq + intro f + have key := congr_fun (congrArg FunLike.coe hμν) f + apply (ENNReal.toNNReal_eq_toNNReal_iff' ?_ ?_).mp key + · exact (lintegral_lt_top_of_nnreal μ f).ne + · exact (lintegral_lt_top_of_nnreal ν f).ne + +variable (Ω) + +lemma embedding_toWeakDualBCNN : + Embedding (toWeakDualBCNN : FiniteMeasure Ω → WeakDual ℝ≥0 (Ω →ᵇ ℝ≥0)) where + induced := rfl + inj := injective_toWeakDualBCNN + +/-- On topological spaces where indicators of closed sets have decreasing approximating sequences of +continuous functions (`HasOuterApproxClosed`), the topology of weak convergence of finite Borel +measures is Hausdorff (`T2Space`). -/ +instance t2Space : T2Space (FiniteMeasure Ω) := + Embedding.t2Space (embedding_toWeakDualBCNN Ω) + +end Hausdorff -- section + end FiniteMeasure -- section @@ -553,35 +604,13 @@ This section is about bounded convergence theorems for finite measures. variable {Ω : Type*} [MeasurableSpace Ω] [TopologicalSpace Ω] [OpensMeasurableSpace Ω] -/-- A bounded convergence theorem for a finite measure: -If bounded continuous non-negative functions are uniformly bounded by a constant and tend to a -limit, then their integrals against the finite measure tend to the integral of the limit. -This formulation assumes: - * the functions tend to a limit along a countably generated filter; - * the limit is in the almost everywhere sense; - * boundedness holds almost everywhere; - * integration is `MeasureTheory.lintegral`, i.e., the functions and their integrals are - `ℝ≥0∞`-valued. --/ -theorem tendsto_lintegral_nn_filter_of_le_const {ι : Type*} {L : Filter ι} [L.IsCountablyGenerated] - (μ : Measure Ω) [IsFiniteMeasure μ] {fs : ι → Ω →ᵇ ℝ≥0} {c : ℝ≥0} - (fs_le_const : ∀ᶠ i in L, ∀ᵐ ω : Ω ∂μ, fs i ω ≤ c) {f : Ω → ℝ≥0} - (fs_lim : ∀ᵐ ω : Ω ∂μ, Tendsto (fun i => fs i ω) L (𝓝 (f ω))) : - Tendsto (fun i => ∫⁻ ω, fs i ω ∂μ) L (𝓝 (∫⁻ ω, f ω ∂μ)) := by - refine tendsto_lintegral_filter_of_dominated_convergence (fun _ => c) - (eventually_of_forall fun i => (ENNReal.continuous_coe.comp (fs i).continuous).measurable) ?_ - (@lintegral_const_lt_top _ _ μ _ _ (@ENNReal.coe_ne_top c)).ne ?_ - · simpa only [Function.comp_apply, ENNReal.coe_le_coe] using fs_le_const - · simpa only [Function.comp_apply, ENNReal.tendsto_coe] using fs_lim -#align measure_theory.finite_measure.tendsto_lintegral_nn_filter_of_le_const MeasureTheory.FiniteMeasure.tendsto_lintegral_nn_filter_of_le_const - /-- A bounded convergence theorem for a finite measure: If a sequence of bounded continuous non-negative functions are uniformly bounded by a constant and tend pointwise to a limit, then their integrals (`MeasureTheory.lintegral`) against the finite measure tend to the integral of the limit. A related result with more general assumptions is -`MeasureTheory.FiniteMeasure.tendsto_lintegral_nn_filter_of_le_const`. +`MeasureTheory.tendsto_lintegral_nn_filter_of_le_const`. -/ theorem tendsto_lintegral_nn_of_le_const (μ : FiniteMeasure Ω) {fs : ℕ → Ω →ᵇ ℝ≥0} {c : ℝ≥0} (fs_le_const : ∀ n ω, fs n ω ≤ c) {f : Ω → ℝ≥0} diff --git a/Mathlib/MeasureTheory/Measure/Haar/InnerProductSpace.lean b/Mathlib/MeasureTheory/Measure/Haar/InnerProductSpace.lean index be0a2773272d5..a8394cb2f66b3 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/InnerProductSpace.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/InnerProductSpace.lean @@ -139,7 +139,7 @@ theorem volume_ball (x : EuclideanSpace ℝ (Fin 2)) (r : ℝ) : have : MeasurePreserving (_ : ℝ × ℝ ≃ᵐ EuclideanSpace ℝ (Fin 2)) := MeasurePreserving.trans (volume_preserving_finTwoArrow ℝ).symm (volume_preserving_measurableEquiv (Fin 2)).symm - rw [←this.measure_preimage_emb (MeasurableEquiv.measurableEmbedding _), + rw [← this.measure_preimage_emb (MeasurableEquiv.measurableEmbedding _), ball_zero_eq _ zero_le_one, preimage_setOf_eq] simp only [MeasurableEquiv.finTwoArrow_symm_apply, Fin.sum_univ_two, preimage_setOf_eq, Fin.cons_zero, Fin.cons_one, one_pow, Function.comp_apply, coe_measurableEquiv_symm, diff --git a/Mathlib/MeasureTheory/Measure/Haar/OfBasis.lean b/Mathlib/MeasureTheory/Measure/Haar/OfBasis.lean index 88633e943f415..ed9feb58563ec 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/OfBasis.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/OfBasis.lean @@ -54,7 +54,7 @@ theorem image_parallelepiped (f : E →ₗ[ℝ] F) (v : ι → E) : f '' parallelepiped v = parallelepiped (f ∘ v) := by simp only [parallelepiped, ← image_comp] congr 1 with t - simp only [Function.comp_apply, map_sum, LinearMap.map_smulₛₗ, RingHom.id_apply] + simp only [Function.comp_apply, _root_.map_sum, LinearMap.map_smulₛₗ, RingHom.id_apply] #align image_parallelepiped image_parallelepiped /-- Reindexing a family of vectors does not change their parallelepiped. -/ diff --git a/Mathlib/MeasureTheory/Measure/HasOuterApproxClosed.lean b/Mathlib/MeasureTheory/Measure/HasOuterApproxClosed.lean new file mode 100644 index 0000000000000..eb406ed3b0f3d --- /dev/null +++ b/Mathlib/MeasureTheory/Measure/HasOuterApproxClosed.lean @@ -0,0 +1,238 @@ +/- +Copyright (c) 2022 Kalle Kytölä. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kalle Kytölä +-/ +import Mathlib.MeasureTheory.Integral.Lebesgue +import Mathlib.Topology.MetricSpace.ThickenedIndicator + +/-! +# Spaces where indicators of closed sets have decreasing approximations by continuous functions + +In this file we define a typeclass `HasOuterApproxClosed` for topological spaces in which indicator +functions of closed sets have sequences of bounded continuous functions approximating them from +above. All pseudo-emetrizable spaces have this property, see `instHasOuterApproxClosed`. + +In spaces with the `HasOuterApproxClosed` property, finite Borel measures are uniquely characterized +by the integrals of bounded continuous functions. Also weak convergence of finite measures and +convergence in distribution for random variables behave somewhat well in spaces with this property. + +## Main definitions + + * `HasOuterApproxClosed`: the typeclass for topological spaces in which indicator functions of + closed sets have sequences of bounded continuous functions approximating them. + * `IsClosed.apprSeq`: a (non-constructive) choice of an approximating sequence to the indicator + function of a closed set. + +## Main results + + * `instHasOuterApproxClosed`: Any pseudo-emetrizable space has the property `HasOuterApproxClosed`. + * `tendsto_lintegral_apprSeq`: The integrals of the approximating functions to the indicator of a + closed set tend to the measure of the set. + * `ext_of_forall_lintegral_eq_of_IsFiniteMeasure`: Two finite measures are equal if the integrals + of all bounded continuous functions with respect to both agree. + +-/ + +open MeasureTheory Topology Metric Filter Set ENNReal NNReal +open scoped Topology ENNReal NNReal BoundedContinuousFunction + +section auxiliary + +namespace MeasureTheory + +variable {Ω : Type*} [TopologicalSpace Ω] [MeasurableSpace Ω] [OpensMeasurableSpace Ω] + +/-- A bounded convergence theorem for a finite measure: +If bounded continuous non-negative functions are uniformly bounded by a constant and tend to a +limit, then their integrals against the finite measure tend to the integral of the limit. +This formulation assumes: + * the functions tend to a limit along a countably generated filter; + * the limit is in the almost everywhere sense; + * boundedness holds almost everywhere; + * integration is `MeasureTheory.lintegral`, i.e., the functions and their integrals are + `ℝ≥0∞`-valued. +-/ +theorem tendsto_lintegral_nn_filter_of_le_const {ι : Type*} {L : Filter ι} [L.IsCountablyGenerated] + (μ : Measure Ω) [IsFiniteMeasure μ] {fs : ι → Ω →ᵇ ℝ≥0} {c : ℝ≥0} + (fs_le_const : ∀ᶠ i in L, ∀ᵐ ω : Ω ∂μ, fs i ω ≤ c) {f : Ω → ℝ≥0} + (fs_lim : ∀ᵐ ω : Ω ∂μ, Tendsto (fun i ↦ fs i ω) L (𝓝 (f ω))) : + Tendsto (fun i ↦ ∫⁻ ω, fs i ω ∂μ) L (𝓝 (∫⁻ ω, f ω ∂μ)) := by + refine tendsto_lintegral_filter_of_dominated_convergence (fun _ ↦ c) + (eventually_of_forall fun i ↦ (ENNReal.continuous_coe.comp (fs i).continuous).measurable) ?_ + (@lintegral_const_lt_top _ _ μ _ _ (@ENNReal.coe_ne_top c)).ne ?_ + · simpa only [Function.comp_apply, ENNReal.coe_le_coe] using fs_le_const + · simpa only [Function.comp_apply, ENNReal.tendsto_coe] using fs_lim +#align measure_theory.finite_measure.tendsto_lintegral_nn_filter_of_le_const MeasureTheory.tendsto_lintegral_nn_filter_of_le_const + +/-- If bounded continuous functions tend to the indicator of a measurable set and are +uniformly bounded, then their integrals against a finite measure tend to the measure of the set. +This formulation assumes: + * the functions tend to a limit along a countably generated filter; + * the limit is in the almost everywhere sense; + * boundedness holds almost everywhere. +-/ +theorem measure_of_cont_bdd_of_tendsto_filter_indicator {ι : Type*} {L : Filter ι} + [L.IsCountablyGenerated] [TopologicalSpace Ω] [OpensMeasurableSpace Ω] (μ : Measure Ω) + [IsFiniteMeasure μ] {c : ℝ≥0} {E : Set Ω} (E_mble : MeasurableSet E) (fs : ι → Ω →ᵇ ℝ≥0) + (fs_bdd : ∀ᶠ i in L, ∀ᵐ ω : Ω ∂μ, fs i ω ≤ c) + (fs_lim : ∀ᵐ ω ∂μ, Tendsto (fun i ↦ fs i ω) L (𝓝 (indicator E (fun _ ↦ (1 : ℝ≥0)) ω))) : + Tendsto (fun n ↦ lintegral μ fun ω ↦ fs n ω) L (𝓝 (μ E)) := by + convert tendsto_lintegral_nn_filter_of_le_const μ fs_bdd fs_lim + have aux : ∀ ω, indicator E (fun _ ↦ (1 : ℝ≥0∞)) ω = ↑(indicator E (fun _ ↦ (1 : ℝ≥0)) ω) := + fun ω ↦ by simp only [ENNReal.coe_indicator, ENNReal.coe_one] + simp_rw [← aux, lintegral_indicator _ E_mble] + simp only [lintegral_one, Measure.restrict_apply, MeasurableSet.univ, univ_inter] +#align measure_theory.measure_of_cont_bdd_of_tendsto_filter_indicator MeasureTheory.measure_of_cont_bdd_of_tendsto_filter_indicator + +/-- If a sequence of bounded continuous functions tends to the indicator of a measurable set and +the functions are uniformly bounded, then their integrals against a finite measure tend to the +measure of the set. + +A similar result with more general assumptions is +`MeasureTheory.measure_of_cont_bdd_of_tendsto_filter_indicator`. +-/ +theorem measure_of_cont_bdd_of_tendsto_indicator [OpensMeasurableSpace Ω] + (μ : Measure Ω) [IsFiniteMeasure μ] {c : ℝ≥0} {E : Set Ω} (E_mble : MeasurableSet E) + (fs : ℕ → Ω →ᵇ ℝ≥0) (fs_bdd : ∀ n ω, fs n ω ≤ c) + (fs_lim : Tendsto (fun n ω ↦ fs n ω) atTop (𝓝 (indicator E fun _ ↦ (1 : ℝ≥0)))) : + Tendsto (fun n ↦ lintegral μ fun ω ↦ fs n ω) atTop (𝓝 (μ E)) := by + have fs_lim' : + ∀ ω, Tendsto (fun n : ℕ ↦ (fs n ω : ℝ≥0)) atTop (𝓝 (indicator E (fun _ ↦ (1 : ℝ≥0)) ω)) := by + rw [tendsto_pi_nhds] at fs_lim + exact fun ω ↦ fs_lim ω + apply measure_of_cont_bdd_of_tendsto_filter_indicator μ E_mble fs + (eventually_of_forall fun n ↦ eventually_of_forall (fs_bdd n)) (eventually_of_forall fs_lim') +#align measure_theory.measure_of_cont_bdd_of_tendsto_indicator MeasureTheory.measure_of_cont_bdd_of_tendsto_indicator + +/-- The integrals of thickened indicators of a closed set against a finite measure tend to the +measure of the closed set if the thickening radii tend to zero. -/ +theorem tendsto_lintegral_thickenedIndicator_of_isClosed {Ω : Type*} [MeasurableSpace Ω] + [PseudoEMetricSpace Ω] [OpensMeasurableSpace Ω] (μ : Measure Ω) [IsFiniteMeasure μ] {F : Set Ω} + (F_closed : IsClosed F) {δs : ℕ → ℝ} (δs_pos : ∀ n, 0 < δs n) + (δs_lim : Tendsto δs atTop (𝓝 0)) : + Tendsto (fun n ↦ lintegral μ fun ω ↦ (thickenedIndicator (δs_pos n) F ω : ℝ≥0∞)) atTop + (𝓝 (μ F)) := by + apply measure_of_cont_bdd_of_tendsto_indicator μ F_closed.measurableSet + (fun n ↦ thickenedIndicator (δs_pos n) F) fun n ω ↦ thickenedIndicator_le_one (δs_pos n) F ω + have key := thickenedIndicator_tendsto_indicator_closure δs_pos δs_lim F + rwa [F_closed.closure_eq] at key +#align measure_theory.tendsto_lintegral_thickened_indicator_of_is_closed MeasureTheory.tendsto_lintegral_thickenedIndicator_of_isClosed + +end MeasureTheory -- namespace + +end auxiliary -- section + +section HasOuterApproxClosed + +/-- A type class for topological spaces in which the indicator functions of closed sets can be +approximated pointwise from above by a sequence of bounded continuous functions. -/ +class HasOuterApproxClosed (X : Type*) [TopologicalSpace X] : Prop where + exAppr : ∀ (F : Set X), IsClosed F → ∃ (fseq : ℕ → (X →ᵇ ℝ≥0)), + (∀ n x, fseq n x ≤ 1) ∧ (∀ n x, x ∈ F → 1 ≤ fseq n x) ∧ + Tendsto (fun n : ℕ ↦ (fun x ↦ fseq n x)) atTop (𝓝 (indicator F fun _ ↦ (1 : ℝ≥0))) + +namespace HasOuterApproxClosed + +variable {X : Type*} [TopologicalSpace X] [HasOuterApproxClosed X] +variable {F : Set X} (hF : IsClosed F) + +/-- A sequence of continuous functions `X → [0,1]` tending to the indicator of a closed set. -/ +noncomputable def _root_.IsClosed.apprSeq : ℕ → (X →ᵇ ℝ≥0) := + Exists.choose (HasOuterApproxClosed.exAppr F hF) + +lemma apprSeq_apply_le_one (n : ℕ) (x : X) : + hF.apprSeq n x ≤ 1 := + (Exists.choose_spec (HasOuterApproxClosed.exAppr F hF)).1 n x + +lemma apprSeq_apply_eq_one (n : ℕ) {x : X} (hxF : x ∈ F) : + hF.apprSeq n x = 1 := + le_antisymm (apprSeq_apply_le_one _ _ _) + ((Exists.choose_spec (HasOuterApproxClosed.exAppr F hF)).2.1 n x hxF) + +lemma tendsto_apprSeq : + Tendsto (fun n : ℕ ↦ (fun x ↦ hF.apprSeq n x)) atTop (𝓝 (indicator F fun _ ↦ (1 : ℝ≥0))) := + (Exists.choose_spec (HasOuterApproxClosed.exAppr F hF)).2.2 + +lemma indicator_le_apprSeq (n : ℕ) : + indicator F (fun _ ↦ 1) ≤ hF.apprSeq n := by + intro x + by_cases hxF : x ∈ F + · simp only [hxF, indicator_of_mem, apprSeq_apply_eq_one hF n, le_refl] + · simp only [hxF, not_false_eq_true, indicator_of_not_mem, zero_le] + +/-- The measure of a closed set is at most the integral of any function in a decreasing +approximating sequence to the indicator of the set. -/ +theorem measure_le_lintegral [MeasurableSpace X] [OpensMeasurableSpace X] (μ : Measure X) (n : ℕ) : + μ F ≤ ∫⁻ x, (hF.apprSeq n x : ℝ≥0∞) ∂μ := by + convert_to ∫⁻ x, (F.indicator (fun _ ↦ (1 : ℝ≥0∞))) x ∂μ ≤ ∫⁻ x, hF.apprSeq n x ∂μ + · rw [lintegral_indicator _ hF.measurableSet] + simp only [lintegral_one, MeasurableSet.univ, Measure.restrict_apply, univ_inter] + · apply lintegral_mono + intro x + by_cases hxF : x ∈ F + · simp only [hxF, indicator_of_mem, apprSeq_apply_eq_one hF n hxF, coe_one, le_refl] + · simp only [hxF, not_false_eq_true, indicator_of_not_mem, zero_le] + +/-- The integrals along a decreasing approximating sequence to the indicator of a closed set +tend to the measure of the closed set. -/ +lemma tendsto_lintegral_apprSeq [MeasurableSpace X] [OpensMeasurableSpace X] + (μ : Measure X) [IsFiniteMeasure μ] : + Tendsto (fun n ↦ ∫⁻ x, hF.apprSeq n x ∂μ) atTop (𝓝 ((μ : Measure X) F)) := + measure_of_cont_bdd_of_tendsto_indicator μ hF.measurableSet hF.apprSeq + (apprSeq_apply_le_one hF) (tendsto_apprSeq hF) + +end HasOuterApproxClosed --namespace + +noncomputable instance (X : Type*) [TopologicalSpace X] + [TopologicalSpace.PseudoMetrizableSpace X] : HasOuterApproxClosed X := by + letI : PseudoMetricSpace X := TopologicalSpace.pseudoMetrizableSpacePseudoMetric X + refine ⟨fun F hF ↦ ?_⟩ + · use fun n ↦ thickenedIndicator (δ := (1 : ℝ) / (n + 1)) Nat.one_div_pos_of_nat F + refine ⟨?_, ⟨?_, ?_⟩⟩ + · exact fun n x ↦ thickenedIndicator_le_one Nat.one_div_pos_of_nat F x + · exact fun n x hxF ↦ one_le_thickenedIndicator_apply X Nat.one_div_pos_of_nat hxF + · have key := thickenedIndicator_tendsto_indicator_closure + (δseq := fun (n : ℕ) ↦ (1 : ℝ) / (n + 1)) + (fun _ ↦ Nat.one_div_pos_of_nat) tendsto_one_div_add_atTop_nhds_0_nat F + rw [tendsto_pi_nhds] at * + intro x + nth_rw 2 [←IsClosed.closure_eq hF] + exact key x + +namespace MeasureTheory + +/-- Two finite measures give equal values to all closed sets if the integrals of all bounded +continuous functions with respect to the two measures agree. -/ +theorem measure_isClosed_eq_of_forall_lintegral_eq_of_isFiniteMeasure {Ω : Type*} + [MeasurableSpace Ω] [TopologicalSpace Ω] [HasOuterApproxClosed Ω] + [OpensMeasurableSpace Ω] {μ ν : Measure Ω} [IsFiniteMeasure μ] + (h : ∀ (f : Ω →ᵇ ℝ≥0), ∫⁻ x, f x ∂μ = ∫⁻ x, f x ∂ν) {F : Set Ω} (F_closed : IsClosed F) : + μ F = ν F := by + have ν_finite : IsFiniteMeasure ν := by + have whole := h 1 + simp only [BoundedContinuousFunction.coe_one, Pi.one_apply, coe_one, lintegral_const, one_mul] + at whole + refine ⟨by simpa [← whole] using IsFiniteMeasure.measure_univ_lt_top⟩ + have obs_μ := HasOuterApproxClosed.tendsto_lintegral_apprSeq F_closed μ + have obs_ν := HasOuterApproxClosed.tendsto_lintegral_apprSeq F_closed ν + simp_rw [h] at obs_μ + exact tendsto_nhds_unique obs_μ obs_ν + +/-- Two finite Borel measures are equal if the integrals of all bounded continuous functions with +respect to both agree. -/ +theorem ext_of_forall_lintegral_eq_of_IsFiniteMeasure {Ω : Type*} + [MeasurableSpace Ω] [TopologicalSpace Ω] [HasOuterApproxClosed Ω] + [BorelSpace Ω] {μ ν : Measure Ω} [IsFiniteMeasure μ] + (h : ∀ (f : Ω →ᵇ ℝ≥0), ∫⁻ x, f x ∂μ = ∫⁻ x, f x ∂ν) : + μ = ν := by + have key := @measure_isClosed_eq_of_forall_lintegral_eq_of_isFiniteMeasure Ω _ _ _ _ μ ν _ h + apply ext_of_generate_finite _ ?_ isPiSystem_isClosed + · exact fun F F_closed ↦ key F_closed + · exact key isClosed_univ + · rw [BorelSpace.measurable_eq (α := Ω), borel_eq_generateFrom_isClosed] + rfl + +end MeasureTheory -- namespace + +end HasOuterApproxClosed -- section diff --git a/Mathlib/MeasureTheory/Measure/Hausdorff.lean b/Mathlib/MeasureTheory/Measure/Hausdorff.lean index 2faecad54a3ec..307dbe3b01522 100644 --- a/Mathlib/MeasureTheory/Measure/Hausdorff.lean +++ b/Mathlib/MeasureTheory/Measure/Hausdorff.lean @@ -617,7 +617,7 @@ theorem hausdorffMeasure_le_liminf_sum {β : Type*} {ι : β → Type*} [∀ n, /-- If `d₁ < d₂`, then for any set `s` we have either `μH[d₂] s = 0`, or `μH[d₁] s = ∞`. -/ theorem hausdorffMeasure_zero_or_top {d₁ d₂ : ℝ} (h : d₁ < d₂) (s : Set X) : μH[d₂] s = 0 ∨ μH[d₁] s = ∞ := by - by_contra' H + by_contra! H suffices ∀ c : ℝ≥0, c ≠ 0 → μH[d₂] s ≤ c * μH[d₁] s by rcases ENNReal.exists_nnreal_pos_mul_lt H.2 H.1 with ⟨c, hc0, hc⟩ exact hc.not_le (this c (pos_iff_ne_zero.1 hc0)) diff --git a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean index 0186e7b67d94f..199c48643ce47 100644 --- a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean +++ b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean @@ -452,7 +452,7 @@ theorem measure_iUnion_eq_iSup [Countable ι] {s : ι → Set α} (hd : Directed generalize ht : Function.extend Encodable.encode s ⊥ = t replace hd : Directed (· ⊆ ·) t := ht ▸ hd.extend_bot Encodable.encode_injective suffices μ (⋃ n, t n) = ⨆ n, μ (t n) by - simp only [← ht, Encodable.encode_injective.apply_extend μ, ← iSup_eq_iUnion, + simp only [← ht, Function.apply_extend μ, ← iSup_eq_iUnion, iSup_extend_bot Encodable.encode_injective, (· ∘ ·), Pi.bot_apply, bot_eq_empty, measure_empty] at this exact this.trans (iSup_extend_bot Encodable.encode_injective _) @@ -478,7 +478,6 @@ theorem measure_iUnion_eq_iSup [Countable ι] {s : ι → Set α} (hd : Directed _ = μ (⋃ n ∈ I, t n) := (measure_biUnion_toMeasurable I.countable_toSet _) _ ≤ μ (t N) := (measure_mono (iUnion₂_subset hN)) _ ≤ ⨆ n, μ (t n) := le_iSup (μ ∘ t) N - #align measure_theory.measure_Union_eq_supr MeasureTheory.measure_iUnion_eq_iSup theorem measure_biUnion_eq_iSup {s : ι → Set α} {t : Set ι} (ht : t.Countable) @@ -872,7 +871,7 @@ def coeAddHom {_ : MeasurableSpace α} : Measure α →+ Set α → ℝ≥0∞ w @[simp] theorem coe_finset_sum {_m : MeasurableSpace α} (I : Finset ι) (μ : ι → Measure α) : - ⇑(∑ i in I, μ i) = ∑ i in I, ⇑(μ i) := coeAddHom.map_sum μ I + ⇑(∑ i in I, μ i) = ∑ i in I, ⇑(μ i) := map_sum coeAddHom μ I #align measure_theory.measure.coe_finset_sum MeasureTheory.Measure.coe_finset_sum theorem finset_sum_apply {m : MeasurableSpace α} (I : Finset ι) (μ : ι → Measure α) (s : Set α) : @@ -1407,7 +1406,8 @@ theorem le_sum_apply (f : ι → Measure α) (s : Set α) : ∑' i, f i s ≤ su #align measure_theory.measure.le_sum_apply MeasureTheory.Measure.le_sum_apply @[simp] -theorem sum_apply (f : ι → Measure α) {s : Set α} (hs : MeasurableSet s) : sum f s = ∑' i, f i s := +theorem sum_apply (f : ι → Measure α) {s : Set α} (hs : MeasurableSet s) : + sum f s = ∑' i, f i s := toMeasure_apply _ _ hs #align measure_theory.measure.sum_apply MeasureTheory.Measure.sum_apply @@ -1420,6 +1420,23 @@ theorem sum_apply₀ (f : ι → Measure α) {s : Set α} (hs : NullMeasurableSe _ = ∑' i, f i t := sum_apply _ t_meas _ ≤ ∑' i, f i s := ENNReal.tsum_le_tsum (fun i ↦ measure_mono ts) +/-! For the next theorem, the countability assumption is necessary. For a counterexample, consider +an uncountable space, with a distinguished point `x₀`, and the sigma-algebra made of countable sets +not containing `x₀`, and their complements. All points but `x₀` are measurable. +Consider the sum of the Dirac masses at points different from `x₀`, and `s = x₀`. For any Dirac mass +`δ_x`, we have `δ_x (x₀) = 0`, so `∑' x, δ_x (x₀) = 0`. On the other hand, the measure `sum δ_x` +gives mass one to each point different from `x₀`, so it gives infinite mass to any measurable set +containing `x₀` (as such a set is uncountable), and by outer regularity one get `sum δ_x {x₀} = ∞`. +-/ +theorem sum_apply_of_countable [Countable ι] (f : ι → Measure α) (s : Set α) : + sum f s = ∑' i, f i s := by + apply le_antisymm ?_ (le_sum_apply _ _) + rcases exists_measurable_superset_forall_eq f s with ⟨t, hst, htm, ht⟩ + calc + sum f s ≤ sum f t := measure_mono hst + _ = ∑' i, f i t := sum_apply _ htm + _ = ∑' i, f i s := by simp [ht] + theorem le_sum (μ : ι → Measure α) (i : ι) : μ i ≤ sum μ := fun s hs => by simpa only [sum_apply μ hs] using ENNReal.le_tsum i #align measure_theory.measure.le_sum MeasureTheory.Measure.le_sum @@ -1441,6 +1458,11 @@ theorem sum_apply_eq_zero' {μ : ι → Measure α} {s : Set α} (hs : Measurabl sum μ s = 0 ↔ ∀ i, μ i s = 0 := by simp [hs] #align measure_theory.measure.sum_apply_eq_zero' MeasureTheory.Measure.sum_apply_eq_zero' +theorem sum_sum {ι' : Type*} (μ : ι → ι' → Measure α) : + (sum fun n => sum (μ n)) = sum (fun (p : ι × ι') ↦ μ p.1 p.2) := by + ext1 s hs + simp [sum_apply _ hs, ENNReal.tsum_prod'] + theorem sum_comm {ι' : Type*} (μ : ι → ι' → Measure α) : (sum fun n => sum (μ n)) = sum fun m => sum fun n => μ n m := by ext1 s hs @@ -1502,12 +1524,21 @@ theorem sum_congr {μ ν : ℕ → Measure α} (h : ∀ n, μ n = ν n) : sum μ congr_arg sum (funext h) #align measure_theory.measure.sum_congr MeasureTheory.Measure.sum_congr -theorem sum_add_sum (μ ν : ℕ → Measure α) : sum μ + sum ν = sum fun n => μ n + ν n := by +theorem sum_add_sum {ι : Type*} (μ ν : ι → Measure α) : sum μ + sum ν = sum fun n => μ n + ν n := by ext1 s hs simp only [add_apply, sum_apply _ hs, Pi.add_apply, coe_add, tsum_add ENNReal.summable ENNReal.summable] #align measure_theory.measure.sum_add_sum MeasureTheory.Measure.sum_add_sum +@[simp] lemma sum_comp_equiv {ι ι' : Type*} (e : ι' ≃ ι) (m : ι → Measure α) : + sum (m ∘ e) = sum m := by + ext s hs + simpa [hs, sum_apply] using e.tsum_eq (fun n ↦ m n s) + +@[simp] lemma sum_extend_zero {ι ι' : Type*} {f : ι → ι'} (hf : Injective f) (m : ι → Measure α) : + sum (Function.extend f m 0) = sum m := by + ext s hs + simp [*, Function.apply_extend (fun μ : Measure α ↦ μ s)] end Sum /-! ### Absolute continuity -/ diff --git a/Mathlib/MeasureTheory/Measure/Portmanteau.lean b/Mathlib/MeasureTheory/Measure/Portmanteau.lean index 20e1b1dc848f0..fc12e03bf071b 100644 --- a/Mathlib/MeasureTheory/Measure/Portmanteau.lean +++ b/Mathlib/MeasureTheory/Measure/Portmanteau.lean @@ -274,98 +274,37 @@ Combining with a earlier proven implications, we get that (T) implies also both -/ -variable {Ω : Type*} [MeasurableSpace Ω] - -/-- If bounded continuous functions tend to the indicator of a measurable set and are -uniformly bounded, then their integrals against a finite measure tend to the measure of the set. -This formulation assumes: - * the functions tend to a limit along a countably generated filter; - * the limit is in the almost everywhere sense; - * boundedness holds almost everywhere. --/ -theorem measure_of_cont_bdd_of_tendsto_filter_indicator {ι : Type*} {L : Filter ι} - [L.IsCountablyGenerated] [TopologicalSpace Ω] [OpensMeasurableSpace Ω] (μ : Measure Ω) - [IsFiniteMeasure μ] {c : ℝ≥0} {E : Set Ω} (E_mble : MeasurableSet E) (fs : ι → Ω →ᵇ ℝ≥0) - (fs_bdd : ∀ᶠ i in L, ∀ᵐ ω : Ω ∂μ, fs i ω ≤ c) - (fs_lim : ∀ᵐ ω : Ω ∂μ, Tendsto (fun i : ι => ((⇑) : (Ω →ᵇ ℝ≥0) → Ω → ℝ≥0) (fs i) ω) L - (𝓝 (indicator E (fun _ => (1 : ℝ≥0)) ω))) : - Tendsto (fun n => lintegral μ fun ω => fs n ω) L (𝓝 (μ E)) := by - convert FiniteMeasure.tendsto_lintegral_nn_filter_of_le_const μ fs_bdd fs_lim - have aux : ∀ ω, indicator E (fun _ => (1 : ℝ≥0∞)) ω = ↑(indicator E (fun _ => (1 : ℝ≥0)) ω) := - fun ω => by simp only [ENNReal.coe_indicator, ENNReal.coe_one] - simp_rw [← aux, lintegral_indicator _ E_mble] - simp only [lintegral_one, Measure.restrict_apply, MeasurableSet.univ, univ_inter] -#align measure_theory.measure_of_cont_bdd_of_tendsto_filter_indicator MeasureTheory.measure_of_cont_bdd_of_tendsto_filter_indicator - -/-- If a sequence of bounded continuous functions tends to the indicator of a measurable set and -the functions are uniformly bounded, then their integrals against a finite measure tend to the -measure of the set. - -A similar result with more general assumptions is -`MeasureTheory.measure_of_cont_bdd_of_tendsto_filter_indicator`. --/ -theorem measure_of_cont_bdd_of_tendsto_indicator [TopologicalSpace Ω] [OpensMeasurableSpace Ω] - (μ : Measure Ω) [IsFiniteMeasure μ] {c : ℝ≥0} {E : Set Ω} (E_mble : MeasurableSet E) - (fs : ℕ → Ω →ᵇ ℝ≥0) (fs_bdd : ∀ n ω, fs n ω ≤ c) - (fs_lim : Tendsto (fun n : ℕ => ((⇑) : (Ω →ᵇ ℝ≥0) → Ω → ℝ≥0) (fs n)) atTop - (𝓝 (indicator E fun _ => (1 : ℝ≥0)))) : - Tendsto (fun n => lintegral μ fun ω => fs n ω) atTop (𝓝 (μ E)) := by - have fs_lim' : - ∀ ω, Tendsto (fun n : ℕ => (fs n ω : ℝ≥0)) atTop (𝓝 (indicator E (fun _ => (1 : ℝ≥0)) ω)) := by - rw [tendsto_pi_nhds] at fs_lim - exact fun ω => fs_lim ω - apply measure_of_cont_bdd_of_tendsto_filter_indicator μ E_mble fs - (eventually_of_forall fun n => eventually_of_forall (fs_bdd n)) (eventually_of_forall fs_lim') -#align measure_theory.measure_of_cont_bdd_of_tendsto_indicator MeasureTheory.measure_of_cont_bdd_of_tendsto_indicator - -/-- The integrals of thickened indicators of a closed set against a finite measure tend to the -measure of the closed set if the thickening radii tend to zero. --/ -theorem tendsto_lintegral_thickenedIndicator_of_isClosed {Ω : Type*} [MeasurableSpace Ω] - [PseudoEMetricSpace Ω] [OpensMeasurableSpace Ω] (μ : Measure Ω) [IsFiniteMeasure μ] {F : Set Ω} - (F_closed : IsClosed F) {δs : ℕ → ℝ} (δs_pos : ∀ n, 0 < δs n) - (δs_lim : Tendsto δs atTop (𝓝 0)) : - Tendsto (fun n => lintegral μ fun ω => (thickenedIndicator (δs_pos n) F ω : ℝ≥0∞)) atTop - (𝓝 (μ F)) := by - apply measure_of_cont_bdd_of_tendsto_indicator μ F_closed.measurableSet - (fun n => thickenedIndicator (δs_pos n) F) fun n ω => thickenedIndicator_le_one (δs_pos n) F ω - have key := thickenedIndicator_tendsto_indicator_closure δs_pos δs_lim F - rwa [F_closed.closure_eq] at key -#align measure_theory.tendsto_lintegral_thickened_indicator_of_is_closed MeasureTheory.tendsto_lintegral_thickenedIndicator_of_isClosed - /-- One implication of the portmanteau theorem: Weak convergence of finite measures implies that the limsup of the measures of any closed set is at most the measure of the closed set under the limit measure. -/ theorem FiniteMeasure.limsup_measure_closed_le_of_tendsto {Ω ι : Type*} {L : Filter ι} - [MeasurableSpace Ω] [PseudoEMetricSpace Ω] [OpensMeasurableSpace Ω] {μ : FiniteMeasure Ω} + [MeasurableSpace Ω] [TopologicalSpace Ω] [HasOuterApproxClosed Ω] + [OpensMeasurableSpace Ω] {μ : FiniteMeasure Ω} {μs : ι → FiniteMeasure Ω} (μs_lim : Tendsto μs L (𝓝 μ)) {F : Set Ω} (F_closed : IsClosed F) : (L.limsup fun i => (μs i : Measure Ω) F) ≤ (μ : Measure Ω) F := by rcases L.eq_or_neBot with rfl | hne · simp only [limsup_bot, bot_le] apply ENNReal.le_of_forall_pos_le_add intro ε ε_pos _ - let δs := fun n : ℕ => (1 : ℝ) / (n + 1) - have δs_pos : ∀ n, 0 < δs n := fun n => Nat.one_div_pos_of_nat - have δs_lim : Tendsto δs atTop (𝓝 0) := tendsto_one_div_add_atTop_nhds_0_nat - have key₁ := - tendsto_lintegral_thickenedIndicator_of_isClosed (μ : Measure Ω) F_closed δs_pos δs_lim + let fs := F_closed.apprSeq + have key₁ : Tendsto (fun n ↦ ∫⁻ ω, (fs n ω : ℝ≥0∞) ∂μ) atTop (𝓝 ((μ : Measure Ω) F)) := + HasOuterApproxClosed.tendsto_lintegral_apprSeq F_closed (μ : Measure Ω) have room₁ : (μ : Measure Ω) F < (μ : Measure Ω) F + ε / 2 := by apply ENNReal.lt_add_right (measure_lt_top (μ : Measure Ω) F).ne (ENNReal.div_pos_iff.mpr ⟨(ENNReal.coe_pos.mpr ε_pos).ne.symm, ENNReal.two_ne_top⟩).ne.symm rcases eventually_atTop.mp (eventually_lt_of_tendsto_lt room₁ key₁) with ⟨M, hM⟩ - have key₂ := - FiniteMeasure.tendsto_iff_forall_lintegral_tendsto.mp μs_lim (thickenedIndicator (δs_pos M) F) + have key₂ := FiniteMeasure.tendsto_iff_forall_lintegral_tendsto.mp μs_lim (fs M) have room₂ : - (lintegral (μ : Measure Ω) fun a => thickenedIndicator (δs_pos M) F a) < - (lintegral (μ : Measure Ω) fun a => thickenedIndicator (δs_pos M) F a) + ε / 2 := by + (lintegral (μ : Measure Ω) fun a => fs M a) < + (lintegral (μ : Measure Ω) fun a => fs M a) + ε / 2 := by apply ENNReal.lt_add_right (ne_of_lt ?_) (ENNReal.div_pos_iff.mpr ⟨(ENNReal.coe_pos.mpr ε_pos).ne.symm, ENNReal.two_ne_top⟩).ne.symm apply BoundedContinuousFunction.lintegral_lt_top_of_nnreal have ev_near := Eventually.mono (eventually_lt_of_tendsto_lt room₂ key₂) fun n => le_of_lt - have ev_near' := Eventually.mono ev_near fun n => le_trans - (measure_le_lintegral_thickenedIndicator (μs n : Measure Ω) F_closed.measurableSet (δs_pos M)) + have ev_near' := Eventually.mono ev_near + (fun n ↦ le_trans (HasOuterApproxClosed.measure_le_lintegral F_closed (μs n) M)) apply (Filter.limsup_le_limsup ev_near').trans rw [limsup_const] apply le_trans (add_le_add (hM M rfl.le).le (le_refl (ε / 2 : ℝ≥0∞))) @@ -377,9 +316,10 @@ Weak convergence of probability measures implies that the limsup of the measures set is at most the measure of the closed set under the limit probability measure. -/ theorem ProbabilityMeasure.limsup_measure_closed_le_of_tendsto {Ω ι : Type*} {L : Filter ι} - [MeasurableSpace Ω] [PseudoEMetricSpace Ω] [OpensMeasurableSpace Ω] {μ : ProbabilityMeasure Ω} - {μs : ι → ProbabilityMeasure Ω} (μs_lim : Tendsto μs L (𝓝 μ)) {F : Set Ω} - (F_closed : IsClosed F) : (L.limsup fun i => (μs i : Measure Ω) F) ≤ (μ : Measure Ω) F := by + [MeasurableSpace Ω] [PseudoEMetricSpace Ω] [OpensMeasurableSpace Ω] [HasOuterApproxClosed Ω] + {μ : ProbabilityMeasure Ω} {μs : ι → ProbabilityMeasure Ω} (μs_lim : Tendsto μs L (𝓝 μ)) + {F : Set Ω} (F_closed : IsClosed F) : + (L.limsup fun i => (μs i : Measure Ω) F) ≤ (μ : Measure Ω) F := by apply FiniteMeasure.limsup_measure_closed_le_of_tendsto ((ProbabilityMeasure.tendsto_nhds_iff_toFiniteMeasure_tendsto_nhds L).mp μs_lim) F_closed #align measure_theory.probability_measure.limsup_measure_closed_le_of_tendsto MeasureTheory.ProbabilityMeasure.limsup_measure_closed_le_of_tendsto @@ -389,8 +329,9 @@ Weak convergence of probability measures implies that the liminf of the measures is at least the measure of the open set under the limit probability measure. -/ theorem ProbabilityMeasure.le_liminf_measure_open_of_tendsto {Ω ι : Type*} {L : Filter ι} - [MeasurableSpace Ω] [PseudoEMetricSpace Ω] [OpensMeasurableSpace Ω] {μ : ProbabilityMeasure Ω} - {μs : ι → ProbabilityMeasure Ω} (μs_lim : Tendsto μs L (𝓝 μ)) {G : Set Ω} (G_open : IsOpen G) : + [MeasurableSpace Ω] [PseudoEMetricSpace Ω] [OpensMeasurableSpace Ω] [HasOuterApproxClosed Ω] + {μ : ProbabilityMeasure Ω} {μs : ι → ProbabilityMeasure Ω} (μs_lim : Tendsto μs L (𝓝 μ)) + {G : Set Ω} (G_open : IsOpen G) : (μ : Measure Ω) G ≤ L.liminf fun i => (μs i : Measure Ω) G := haveI h_closeds : ∀ F, IsClosed F → (L.limsup fun i ↦ (μs i : Measure Ω) F) ≤ (μ : Measure Ω) F := fun _ F_closed => ProbabilityMeasure.limsup_measure_closed_le_of_tendsto μs_lim F_closed @@ -400,8 +341,8 @@ theorem ProbabilityMeasure.le_liminf_measure_open_of_tendsto {Ω ι : Type*} {L theorem ProbabilityMeasure.tendsto_measure_of_null_frontier_of_tendsto' {Ω ι : Type*} {L : Filter ι} [MeasurableSpace Ω] [PseudoEMetricSpace Ω] [OpensMeasurableSpace Ω] - {μ : ProbabilityMeasure Ω} {μs : ι → ProbabilityMeasure Ω} (μs_lim : Tendsto μs L (𝓝 μ)) - {E : Set Ω} (E_nullbdry : (μ : Measure Ω) (frontier E) = 0) : + [HasOuterApproxClosed Ω] {μ : ProbabilityMeasure Ω} {μs : ι → ProbabilityMeasure Ω} + (μs_lim : Tendsto μs L (𝓝 μ)) {E : Set Ω} (E_nullbdry : (μ : Measure Ω) (frontier E) = 0) : Tendsto (fun i => (μs i : Measure Ω) E) L (𝓝 ((μ : Measure Ω) E)) := haveI h_opens : ∀ G, IsOpen G → (μ : Measure Ω) G ≤ L.liminf fun i => (μs i : Measure Ω) G := fun _ G_open => ProbabilityMeasure.le_liminf_measure_open_of_tendsto μs_lim G_open @@ -417,9 +358,9 @@ A version with coercions to ordinary `ℝ≥0∞`-valued measures is `MeasureTheory.ProbabilityMeasure.tendsto_measure_of_null_frontier_of_tendsto'`. -/ theorem ProbabilityMeasure.tendsto_measure_of_null_frontier_of_tendsto {Ω ι : Type*} {L : Filter ι} - [MeasurableSpace Ω] [PseudoEMetricSpace Ω] [OpensMeasurableSpace Ω] {μ : ProbabilityMeasure Ω} - {μs : ι → ProbabilityMeasure Ω} (μs_lim : Tendsto μs L (𝓝 μ)) {E : Set Ω} - (E_nullbdry : μ (frontier E) = 0) : Tendsto (fun i => μs i E) L (𝓝 (μ E)) := by + [MeasurableSpace Ω] [PseudoEMetricSpace Ω] [OpensMeasurableSpace Ω] [HasOuterApproxClosed Ω] + {μ : ProbabilityMeasure Ω} {μs : ι → ProbabilityMeasure Ω} (μs_lim : Tendsto μs L (𝓝 μ)) + {E : Set Ω} (E_nullbdry : μ (frontier E) = 0) : Tendsto (fun i => μs i E) L (𝓝 (μ E)) := by have E_nullbdry' : (μ : Measure Ω) (frontier E) = 0 := by rw [← ProbabilityMeasure.ennreal_coeFn_eq_coeFn_toMeasure, E_nullbdry, ENNReal.coe_zero] have key := ProbabilityMeasure.tendsto_measure_of_null_frontier_of_tendsto' μs_lim E_nullbdry' diff --git a/Mathlib/MeasureTheory/Measure/ProbabilityMeasure.lean b/Mathlib/MeasureTheory/Measure/ProbabilityMeasure.lean index b730b3cb142d7..6a7e4de483338 100644 --- a/Mathlib/MeasureTheory/Measure/ProbabilityMeasure.lean +++ b/Mathlib/MeasureTheory/Measure/ProbabilityMeasure.lean @@ -46,6 +46,9 @@ The main definitions are * `MeasureTheory.ProbabilityMeasure.continuous_map`: For a continuous function `f : Ω → Ω'`, the push-forward of probability measures `f* : ProbabilityMeasure Ω → ProbabilityMeasure Ω'` is continuous. + * `MeasureTheory.ProbabilityMeasure.t2Space`: The topology of convergence in distribution is + Hausdorff on Borel spaces where indicators of closed sets have continuous decreasing + approximating sequences (in particular on any pseudo-metrizable spaces). TODO: * Probability measures form a convex space. @@ -214,6 +217,8 @@ theorem toFiniteMeasure_nonzero (μ : ProbabilityMeasure Ω) : μ.toFiniteMeasur exact one_ne_zero #align measure_theory.probability_measure.to_finite_measure_nonzero MeasureTheory.ProbabilityMeasure.toFiniteMeasure_nonzero +section convergence_in_distribution + variable [TopologicalSpace Ω] [OpensMeasurableSpace Ω] theorem testAgainstNN_lipschitz (μ : ProbabilityMeasure Ω) : @@ -300,6 +305,22 @@ theorem tendsto_iff_forall_integral_tendsto {γ : Type*} {F : Filter γ} rfl #align measure_theory.probability_measure.tendsto_iff_forall_integral_tendsto MeasureTheory.ProbabilityMeasure.tendsto_iff_forall_integral_tendsto +end convergence_in_distribution -- section + +section Hausdorff + +variable [TopologicalSpace Ω] [HasOuterApproxClosed Ω] [BorelSpace Ω] + +variable (Ω) + +/-- On topological spaces where indicators of closed sets have decreasing approximating sequences of +continuous functions (`HasOuterApproxClosed`), the topology of convergence in distribution of Borel +probability measures is Hausdorff (`T2Space`). -/ +instance t2Space : T2Space (ProbabilityMeasure Ω) := + Embedding.t2Space (toFiniteMeasure_embedding Ω) + +end Hausdorff -- section + end ProbabilityMeasure -- namespace @@ -333,7 +354,7 @@ def normalize : ProbabilityMeasure Ω := rw [FiniteMeasure.toMeasure_smul] simp only [Measure.smul_toOuterMeasure, OuterMeasure.coe_smul, Pi.smul_apply, Measure.nnreal_smul_coe_apply, ne_eq, mass_zero_iff, ENNReal.coe_inv zero, ennreal_mass] - rw [←Ne.def, ←ENNReal.coe_ne_zero, ennreal_mass] at zero + rw [← Ne.def, ← ENNReal.coe_ne_zero, ennreal_mass] at zero exact ENNReal.inv_mul_cancel zero μ.prop.measure_univ_lt_top.ne } #align measure_theory.finite_measure.normalize MeasureTheory.FiniteMeasure.normalize diff --git a/Mathlib/MeasureTheory/Measure/Regular.lean b/Mathlib/MeasureTheory/Measure/Regular.lean index 233537d60b874..b10ae518d1f1a 100644 --- a/Mathlib/MeasureTheory/Measure/Regular.lean +++ b/Mathlib/MeasureTheory/Measure/Regular.lean @@ -60,7 +60,7 @@ satisfying a predicate `q` with respect to sets satisfying a predicate `p` if fo `U ∈ {U | q U}` and a number `r < μ U` there exists `F ⊆ U` such that `p F` and `r < μ F`. There are two main nontrivial results in the development below: -* `InnerRegularWRT.measurableSet_of_open` shows that, for an outer regular measure, inner +* `InnerRegularWRT.measurableSet_of_isOpen` shows that, for an outer regular measure, inner regularity for open sets with respect to compact sets or closed sets implies inner regularity for all measurable sets of finite measure (with respect to compact sets or closed sets respectively). * `InnerRegularWRT.weaklyRegular_of_finite` shows that a finite measure which is inner regular for @@ -455,7 +455,7 @@ variable {p q : Set α → Prop} {U s : Set α} {ε r : ℝ≥0∞} /-- If a measure is inner regular (using closed or compact sets) for open sets, then every measurable set of finite measure can be approximated by a (closed or compact) subset. -/ -theorem measurableSet_of_open [OuterRegular μ] (H : InnerRegularWRT μ p IsOpen) +theorem measurableSet_of_isOpen [OuterRegular μ] (H : InnerRegularWRT μ p IsOpen) (hd : ∀ ⦃s U⦄, p s → IsOpen U → p (s \ U)) : InnerRegularWRT μ p fun s => MeasurableSet s ∧ μ s ≠ ∞ := by rintro s ⟨hs, hμs⟩ r hr @@ -479,7 +479,7 @@ theorem measurableSet_of_open [OuterRegular μ] (H : InnerRegularWRT μ p IsOpen apply add_le_add_right; apply add_le_add_left exact hμU'.le _ = μ (K \ U') + (ε + ε) := add_assoc _ _ _ -#align measure_theory.measure.inner_regular.measurable_set_of_open MeasureTheory.Measure.InnerRegularWRT.measurableSet_of_open +#align measure_theory.measure.inner_regular.measurable_set_of_open MeasureTheory.Measure.InnerRegularWRT.measurableSet_of_isOpen open Finset in /-- In a finite measure space, assume that any open set can be approximated from inside by closed @@ -863,7 +863,7 @@ theorem _root_.IsOpen.measure_eq_iSup_isClosed ⦃U : Set α⦄ (hU : IsOpen U) theorem innerRegular_measurable [WeaklyRegular μ] : InnerRegularWRT μ IsClosed fun s => MeasurableSet s ∧ μ s ≠ ∞ := - WeaklyRegular.innerRegular.measurableSet_of_open (fun _ _ h₁ h₂ ↦ h₁.inter h₂.isClosed_compl) + WeaklyRegular.innerRegular.measurableSet_of_isOpen (fun _ _ h₁ h₂ ↦ h₁.inter h₂.isClosed_compl) #align measure_theory.measure.weakly_regular.inner_regular_measurable MeasureTheory.Measure.WeaklyRegular.innerRegular_measurable /-- If `s` is a measurable set, a weakly regular measure `μ` is finite on `s`, and `ε` is a positive @@ -969,7 +969,7 @@ theorem exists_compact_not_null [Regular μ] : (∃ K, IsCompact K ∧ μ K ≠ compact subset. See also `MeasurableSet.exists_isCompact_lt_add` and `MeasurableSet.exists_lt_isCompact_of_ne_top`. -/ instance (priority := 100) [Regular μ] : InnerRegularCompactLTTop μ := - ⟨Regular.innerRegular.measurableSet_of_open (fun _ _ hs hU ↦ hs.diff hU)⟩ + ⟨Regular.innerRegular.measurableSet_of_isOpen (fun _ _ hs hU ↦ hs.diff hU)⟩ #noalign measure_theory.measure.regular.inner_regular_measurable protected theorem map [BorelSpace α] [MeasurableSpace β] [TopologicalSpace β] diff --git a/Mathlib/MeasureTheory/Measure/Restrict.lean b/Mathlib/MeasureTheory/Measure/Restrict.lean index fa11540917872..8b4146a28fd87 100644 --- a/Mathlib/MeasureTheory/Measure/Restrict.lean +++ b/Mathlib/MeasureTheory/Measure/Restrict.lean @@ -519,6 +519,12 @@ theorem restrict_sum (μ : ι → Measure α) {s : Set α} (hs : MeasurableSet s ext fun t ht => by simp only [sum_apply, restrict_apply, ht, ht.inter hs] #align measure_theory.measure.restrict_sum MeasureTheory.Measure.restrict_sum +@[simp] +theorem restrict_sum_of_countable [Countable ι] (μ : ι → Measure α) (s : Set α) : + (sum μ).restrict s = sum fun i => (μ i).restrict s := by + ext t ht + simp_rw [sum_apply _ ht, restrict_apply ht, sum_apply_of_countable] + lemma AbsolutelyContinuous.restrict (h : μ ≪ ν) (s : Set α) : μ.restrict s ≪ ν.restrict s := by refine Measure.AbsolutelyContinuous.mk (fun t ht htν ↦ ?_) rw [restrict_apply ht] at htν ⊢ diff --git a/Mathlib/MeasureTheory/Measure/Typeclasses.lean b/Mathlib/MeasureTheory/Measure/Typeclasses.lean index 70c8060fd2c9d..f957f76d25591 100644 --- a/Mathlib/MeasureTheory/Measure/Typeclasses.lean +++ b/Mathlib/MeasureTheory/Measure/Typeclasses.lean @@ -494,7 +494,27 @@ instance isFiniteMeasure_sFiniteSeq [h : SFinite μ] (n : ℕ) : IsFiniteMeasure lemma sum_sFiniteSeq (μ : Measure α) [h : SFinite μ] : sum (sFiniteSeq μ) = μ := h.1.choose_spec.2.symm +/-- A countable sum of finite measures is s-finite. +This lemma is superseeded by the instance below. -/ +lemma sfinite_sum_of_countable {ι : Type*} [Countable ι] + (m : ι → Measure α) [∀ n, IsFiniteMeasure (m n)] : SFinite (Measure.sum m) := by + classical + obtain ⟨f, hf⟩ : ∃ f : ι → ℕ, Function.Injective f := Countable.exists_injective_nat ι + refine ⟨_, fun n ↦ ?_, (sum_extend_zero hf m).symm⟩ + rcases em (n ∈ range f) with ⟨i, rfl⟩ | hn + · rw [hf.extend_apply] + infer_instance + · rw [Function.extend_apply' _ _ _ hn, Pi.zero_apply] + infer_instance + +instance {ι : Type*} [Countable ι] (m : ι → Measure α) [∀ n, SFinite (m n)] : + SFinite (Measure.sum m) := by + change SFinite (Measure.sum (fun i ↦ m i)) + simp_rw [← sum_sFiniteSeq (m _), Measure.sum_sum] + apply sfinite_sum_of_countable + end SFinite + /-- A measure `μ` is called σ-finite if there is a countable collection of sets `{ A i | i ∈ ℕ }` such that `μ (A i) < ∞` and `⋃ i, A i = s`. -/ class SigmaFinite {m0 : MeasurableSpace α} (μ : Measure α) : Prop where @@ -970,7 +990,7 @@ theorem sigmaFinite_bot_iff (μ : @Measure α ⊥) : SigmaFinite μ ↔ IsFinite · rw [h_univ_empty, @measure_empty α ⊥] exact ENNReal.zero_ne_top.lt_top obtain ⟨i, hsi⟩ : ∃ i, s i = Set.univ := by - by_contra' h_not_univ + by_contra! h_not_univ have h_empty : ∀ i, s i = ∅ := by simpa [h_not_univ] using hs_meas simp only [h_empty, iUnion_empty] at hs_univ exact h_univ_empty hs_univ.symm @@ -1093,7 +1113,7 @@ instance isLocallyFiniteMeasureSMulNNReal [TopologicalSpace α] (μ : Measure α protected theorem Measure.isTopologicalBasis_isOpen_lt_top [TopologicalSpace α] (μ : Measure α) [IsLocallyFiniteMeasure μ] : TopologicalSpace.IsTopologicalBasis { s | IsOpen s ∧ μ s < ∞ } := by - refine' TopologicalSpace.isTopologicalBasis_of_open_of_nhds (fun s hs => hs.1) _ + refine' TopologicalSpace.isTopologicalBasis_of_isOpen_of_nhds (fun s hs => hs.1) _ intro x s xs hs rcases μ.exists_isOpen_measure_lt_top x with ⟨v, xv, hv, μv⟩ refine' ⟨v ∩ s, ⟨hv.inter hs, lt_of_le_of_lt _ μv⟩, ⟨xv, xs⟩, inter_subset_right _ _⟩ diff --git a/Mathlib/ModelTheory/Satisfiability.lean b/Mathlib/ModelTheory/Satisfiability.lean index 13fc3262995d1..9a1e8b80f8dcb 100644 --- a/Mathlib/ModelTheory/Satisfiability.lean +++ b/Mathlib/ModelTheory/Satisfiability.lean @@ -683,8 +683,7 @@ theorem Categorical.isComplete (h : κ.Categorical T) (h1 : ℵ₀ ≤ κ) ⟨hS, fun φ => by obtain ⟨_, _⟩ := Theory.exists_model_card_eq ⟨hS.some, hT hS.some⟩ κ h1 h2 rw [Theory.models_sentence_iff, Theory.models_sentence_iff] - by_contra con - push_neg at con + by_contra! con obtain ⟨⟨MF, hMF⟩, MT, hMT⟩ := con rw [Sentence.realize_not, Classical.not_not] at hMT refine' hMF _ diff --git a/Mathlib/NumberTheory/ArithmeticFunction.lean b/Mathlib/NumberTheory/ArithmeticFunction.lean index e753555a175cd..c9137fd9ad289 100644 --- a/Mathlib/NumberTheory/ArithmeticFunction.lean +++ b/Mathlib/NumberTheory/ArithmeticFunction.lean @@ -768,7 +768,7 @@ theorem lcm_apply_mul_gcd_apply [CommMonoidWithZero R] {f : ArithmeticFunction R iterate 4 rw [hf.multiplicative_factorization f (by assumption), Finsupp.prod_of_support_subset _ _ _ (fun _ _ => hfi_zero) (s := (x.primeFactors ⊔ y.primeFactors))] - · rw [←Finset.prod_mul_distrib, ←Finset.prod_mul_distrib] + · rw [← Finset.prod_mul_distrib, ← Finset.prod_mul_distrib] apply Finset.prod_congr rfl intro p _ rcases Nat.le_or_le (x.factorization p) (y.factorization p) with h | h <;> @@ -1207,19 +1207,19 @@ theorem sum_eq_iff_sum_smul_moebius_eq_on [AddCommGroup R] {f g : ℕ → R} let G := fun (n:ℕ) => (∑ i in n.divisors, f i) intro n hn hnP suffices ∑ d in n.divisors, μ (n/d) • G d = f n from by - rw [Nat.sum_divisorsAntidiagonal' (f := fun x y => μ x • g y), ←this, sum_congr rfl] + rw [Nat.sum_divisorsAntidiagonal' (f := fun x y => μ x • g y), ← this, sum_congr rfl] intro d hd - rw [←h d (Nat.pos_of_mem_divisors hd) $ hs d n (Nat.dvd_of_mem_divisors hd) hnP] - rw [←Nat.sum_divisorsAntidiagonal' (f := fun x y => μ x • G y)] + rw [← h d (Nat.pos_of_mem_divisors hd) $ hs d n (Nat.dvd_of_mem_divisors hd) hnP] + rw [← Nat.sum_divisorsAntidiagonal' (f := fun x y => μ x • G y)] apply sum_eq_iff_sum_smul_moebius_eq.mp _ n hn intro _ _; rfl · intro h let F := fun (n:ℕ) => ∑ x : ℕ × ℕ in n.divisorsAntidiagonal, μ x.fst • g x.snd intro n hn hnP suffices ∑ d in n.divisors, F d = g n from by - rw [←this, sum_congr rfl] + rw [← this, sum_congr rfl] intro d hd - rw [←h d (Nat.pos_of_mem_divisors hd) $ hs d n (Nat.dvd_of_mem_divisors hd) hnP] + rw [← h d (Nat.pos_of_mem_divisors hd) $ hs d n (Nat.dvd_of_mem_divisors hd) hnP] apply sum_eq_iff_sum_smul_moebius_eq.mpr _ n hn intro _ _; rfl diff --git a/Mathlib/NumberTheory/Bernoulli.lean b/Mathlib/NumberTheory/Bernoulli.lean index 1df2f17794974..69719ad7e96cd 100644 --- a/Mathlib/NumberTheory/Bernoulli.lean +++ b/Mathlib/NumberTheory/Bernoulli.lean @@ -356,7 +356,7 @@ theorem sum_range_pow (n p : ℕ) : -- key step: a chain of equalities of power series -- porting note: altered proof slightly rw [← mul_right_inj' hexp, mul_comm] - rw [←exp_pow_sum, geom_sum_mul, h_r, ← bernoulliPowerSeries_mul_exp_sub_one, + rw [← exp_pow_sum, geom_sum_mul, h_r, ← bernoulliPowerSeries_mul_exp_sub_one, bernoulliPowerSeries, mul_right_comm] simp only [mul_comm, mul_eq_mul_left_iff, hexp, or_false] refine' Eq.trans (mul_eq_mul_right_iff.mpr _) (Eq.trans h_cauchy _) diff --git a/Mathlib/NumberTheory/ClassNumber/AdmissibleCardPowDegree.lean b/Mathlib/NumberTheory/ClassNumber/AdmissibleCardPowDegree.lean index 0da1ed16ccf90..814f1a5c0c94a 100644 --- a/Mathlib/NumberTheory/ClassNumber/AdmissibleCardPowDegree.lean +++ b/Mathlib/NumberTheory/ClassNumber/AdmissibleCardPowDegree.lean @@ -210,8 +210,7 @@ theorem exists_partition_polynomial_aux (n : ℕ) {ε : ℝ} (hε : 0 < ε) {b : -- but not that `j` is uniquely defined (which is needed to keep the induction going). obtain ⟨j, hj⟩ : ∃ j, ∀ i : Fin n, t' i = j → (cardPowDegree (A 0 % b - A i.succ % b) : ℝ) < cardPowDegree b • ε := by - by_contra hg - push_neg at hg + by_contra! hg obtain ⟨j₀, j₁, j_ne, approx⟩ := exists_approx_polynomial hb hε (Fin.cons (A 0) fun j => A (Fin.succ (Classical.choose (hg j)))) revert j_ne approx diff --git a/Mathlib/NumberTheory/ClassNumber/Finite.lean b/Mathlib/NumberTheory/ClassNumber/Finite.lean index 3116973da911d..19f1281050748 100644 --- a/Mathlib/NumberTheory/ClassNumber/Finite.lean +++ b/Mathlib/NumberTheory/ClassNumber/Finite.lean @@ -69,7 +69,7 @@ noncomputable def normBound : ℤ := theorem normBound_pos : 0 < normBound abv bS := by obtain ⟨i, j, k, hijk⟩ : ∃ i j k, Algebra.leftMulMatrix bS (bS i) j k ≠ 0 := by - by_contra' h + by_contra! h obtain ⟨i⟩ := bS.index_nonempty apply bS.ne_zero i apply diff --git a/Mathlib/NumberTheory/Cyclotomic/Gal.lean b/Mathlib/NumberTheory/Cyclotomic/Gal.lean index b3231e879c764..d68c7a0747447 100644 --- a/Mathlib/NumberTheory/Cyclotomic/Gal.lean +++ b/Mathlib/NumberTheory/Cyclotomic/Gal.lean @@ -136,8 +136,8 @@ noncomputable def autEquivPow : (L ≃ₐ[K] L) ≃* (ZMod n)ˣ := rw [← hζ.val_toRootsOfUnity_coe] simp only [← rootsOfUnity.coe_pow] at key replace key := rootsOfUnity.coe_injective key - rw [pow_eq_pow_iff_modEq, ←Subgroup.orderOf_coe, ← orderOf_units, hζ.val_toRootsOfUnity_coe, ← - (zeta_spec n K L).eq_orderOf, ← ZMod.eq_iff_modEq_nat] at key + rw [pow_eq_pow_iff_modEq, ← Subgroup.orderOf_coe, ← orderOf_units, hζ.val_toRootsOfUnity_coe, + ← (zeta_spec n K L).eq_orderOf, ← ZMod.eq_iff_modEq_nat] at key simp only [ZMod.nat_cast_val, ZMod.cast_id', id.def] at key exact Units.ext key } #align is_cyclotomic_extension.aut_equiv_pow IsCyclotomicExtension.autEquivPow diff --git a/Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean b/Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean index 31f2290895393..9ef15c8d2cbd4 100644 --- a/Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean +++ b/Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean @@ -218,7 +218,7 @@ theorem norm_eq_one [IsDomain L] [IsCyclotomicExtension {n} K L] (hn : n ≠ 2) · rw [h1, one_coe, one_right_iff] at hζ rw [hζ, show 1 = algebraMap K L 1 by simp, Algebra.norm_algebraMap, one_pow] · replace h1 : 2 ≤ n - · by_contra' h + · by_contra! h exact h1 (PNat.eq_one_of_lt_two h) -- Porting note: specyfing the type of `cyclotomic_coeff_zero K h1` was not needed. rw [← hζ.powerBasis_gen K, PowerBasis.norm_gen_eq_coeff_zero_minpoly, hζ.powerBasis_gen K, ← diff --git a/Mathlib/NumberTheory/DiophantineApproximation.lean b/Mathlib/NumberTheory/DiophantineApproximation.lean index 783414893392d..6e421ac4f473f 100644 --- a/Mathlib/NumberTheory/DiophantineApproximation.lean +++ b/Mathlib/NumberTheory/DiophantineApproximation.lean @@ -448,7 +448,7 @@ private theorem aux₂ : 0 < u - ⌊ξ⌋ * v ∧ u - ⌊ξ⌋ * v < v := by -- Porting note: this abused the definitional equality `-1 + 1 = 0` -- refine' (zero_le_mul_right hv₁).mp ((lt_iff_add_one_le (-1 : ℤ) _).mp _) refine' (zero_le_mul_right hv₁).mp ?_ - rw [←sub_one_lt_iff, zero_sub] + rw [← sub_one_lt_iff, zero_sub] replace h := h.1 rw [← lt_sub_iff_add_lt, ← mul_assoc, ← sub_mul] at h exact mod_cast diff --git a/Mathlib/NumberTheory/DirichletCharacter/Basic.lean b/Mathlib/NumberTheory/DirichletCharacter/Basic.lean index 6dcbd5bf14dcb..366efbdc2bfa1 100644 --- a/Mathlib/NumberTheory/DirichletCharacter/Basic.lean +++ b/Mathlib/NumberTheory/DirichletCharacter/Basic.lean @@ -25,7 +25,6 @@ Main definitions: ## TODO -- even, odd - add `lemma unitsMap_surjective {n m : ℕ} (h : n ∣ m) (hm : m ≠ 0) : Function.Surjective (unitsMap h)` and then @@ -225,4 +224,35 @@ lemma primitiveCharacter_one (hn : n ≠ 0) : (isPrimitive_def _).1 (1 : DirichletCharacter R n).primitiveCharacter_isPrimitive, conductor_one hn] +variable {S : Type} [CommRing S] {m : ℕ} (ψ : DirichletCharacter S m) + +/-- A Dirichlet character is odd if its value at -1 is -1. -/ +def Odd : Prop := ψ (-1) = -1 + +/-- A Dirichlet character is even if its value at -1 is 1. -/ +def Even : Prop := ψ (-1) = 1 + +lemma even_or_odd [NoZeroDivisors S] : ψ.Even ∨ ψ.Odd := by + suffices : ψ (-1) ^ 2 = 1 + · convert sq_eq_one_iff.mp this + · rw [← map_pow _, neg_one_sq, map_one] + +lemma Odd.toUnitHom_eval_neg_one (hψ : ψ.Odd) : ψ.toUnitHom (-1) = -1 := by + rw [← Units.eq_iff, MulChar.coe_toUnitHom] + exact hψ + +lemma Even.toUnitHom_eval_neg_one (hψ : ψ.Even) : ψ.toUnitHom (-1) = 1 := by + rw [← Units.eq_iff, MulChar.coe_toUnitHom] + exact hψ + +lemma Odd.eval_neg (x : ZMod m) (hψ : ψ.Odd) : ψ (- x) = - ψ x := by + rw [Odd] at hψ + rw [← neg_one_mul, map_mul] + simp [hψ] + +lemma Even.eval_neg (x : ZMod m) (hψ : ψ.Even) : ψ (- x) = ψ x := by + rw [Even] at hψ + rw [← neg_one_mul, map_mul] + simp [hψ] + end DirichletCharacter diff --git a/Mathlib/NumberTheory/Divisors.lean b/Mathlib/NumberTheory/Divisors.lean index 1c5b80e7a37db..55c241255ab9b 100644 --- a/Mathlib/NumberTheory/Divisors.lean +++ b/Mathlib/NumberTheory/Divisors.lean @@ -318,9 +318,8 @@ theorem perfect_iff_sum_divisors_eq_two_mul (h : 0 < n) : #align nat.perfect_iff_sum_divisors_eq_two_mul Nat.perfect_iff_sum_divisors_eq_two_mul theorem mem_divisors_prime_pow {p : ℕ} (pp : p.Prime) (k : ℕ) {x : ℕ} : - x ∈ divisors (p ^ k) ↔ ∃ (j : ℕ) (_ : j ≤ k), x = p ^ j := by + x ∈ divisors (p ^ k) ↔ ∃ j ≤ k, x = p ^ j := by rw [mem_divisors, Nat.dvd_prime_pow pp, and_iff_left (ne_of_gt (pow_pos pp.pos k))] - simp #align nat.mem_divisors_prime_pow Nat.mem_divisors_prime_pow theorem Prime.divisors {p : ℕ} (pp : p.Prime) : divisors p = {1, p} := by @@ -333,18 +332,11 @@ theorem Prime.properDivisors {p : ℕ} (pp : p.Prime) : properDivisors p = {1} : pp.divisors, pair_comm, erase_insert fun con => pp.ne_one (mem_singleton.1 con)] #align nat.prime.proper_divisors Nat.Prime.properDivisors --- Porting note: Specified pow to Nat.pow theorem divisors_prime_pow {p : ℕ} (pp : p.Prime) (k : ℕ) : - divisors (p ^ k) = (Finset.range (k + 1)).map ⟨Nat.pow p, pow_right_injective pp.two_le⟩ := by + divisors (p ^ k) = (Finset.range (k + 1)).map ⟨(p ^ ·), pow_right_injective pp.two_le⟩ := by ext a - simp only [mem_divisors, mem_map, mem_range, lt_succ_iff, Function.Embedding.coeFn_mk, Nat.pow_eq, - mem_divisors_prime_pow pp k] - have := mem_divisors_prime_pow pp k (x := a) - rw [mem_divisors] at this - rw [this] - refine ⟨?_, ?_⟩ - · intro h; rcases h with ⟨x, hx, hap⟩; use x; tauto - · tauto + rw [mem_divisors_prime_pow pp] + simp [Nat.lt_succ, eq_comm] #align nat.divisors_prime_pow Nat.divisors_prime_pow theorem divisors_injective : Function.Injective divisors := @@ -417,7 +409,7 @@ theorem properDivisors_eq_singleton_one_iff_prime : n.properDivisors = {1} ↔ n have := Nat.le_of_dvd ?_ hdvd · simp [hdvd, this] exact (le_iff_eq_or_lt.mp this).symm - · by_contra' + · by_contra! simp only [nonpos_iff_eq_zero.mp this, this] at h contradiction · exact fun h => Prime.properDivisors h diff --git a/Mathlib/NumberTheory/FLT/Basic.lean b/Mathlib/NumberTheory/FLT/Basic.lean index c9cd11cd9e042..d495d974ff33e 100644 --- a/Mathlib/NumberTheory/FLT/Basic.lean +++ b/Mathlib/NumberTheory/FLT/Basic.lean @@ -64,7 +64,7 @@ lemma fermatLastTheoremWith_nat_int_rat_tfae (n : ℕ) : · refine' h a.natAbs b.natAbs c.natAbs (by positivity) (by positivity) (by positivity) (Int.coe_nat_inj'.1 _) push_cast - simp only [abs_of_neg, neg_pow a, neg_pow b, neg_pow c, ←mul_add, habc, *] + simp only [abs_of_neg, neg_pow a, neg_pow b, neg_pow c, ← mul_add, habc, *] · exact (by positivity : 0 < c ^ n).not_lt $ habc.symm.trans_lt $ add_neg (hn.pow_neg ha) $ hn.pow_neg hb · refine' h b.natAbs c.natAbs a.natAbs (by positivity) (by positivity) (by positivity) @@ -94,7 +94,7 @@ lemma fermatLastTheoremWith_nat_int_rat_tfae (n : ℕ) : simp only [abs_of_pos, habc, *] tfae_have 2 → 3 · rintro h a b c ha hb hc habc - rw [←Rat.num_ne_zero] at ha hb hc + rw [← Rat.num_ne_zero] at ha hb hc have : a.den ≠ 0 := a.den_pos.ne' have : b.den ≠ 0 := b.den_pos.ne' have : c.den ≠ 0 := c.den_pos.ne' @@ -103,7 +103,7 @@ lemma fermatLastTheoremWith_nat_int_rat_tfae (n : ℕ) : have : (a.den * b.den * c.den : ℚ) ^ n ≠ 0 := by positivity refine' Int.cast_injective $ (div_left_inj' this).1 _ push_cast - simp only [add_div, ←div_pow, mul_div_mul_comm, div_self (by positivity : (a.den : ℚ) ≠ 0), + simp only [add_div, ← div_pow, mul_div_mul_comm, div_self (by positivity : (a.den : ℚ) ≠ 0), div_self (by positivity : (b.den : ℚ) ≠ 0), div_self (by positivity : (c.den : ℚ) ≠ 0), one_mul, mul_one, Rat.num_div_den, habc] tfae_have 3 → 1 diff --git a/Mathlib/NumberTheory/LegendreSymbol/AddCharacter.lean b/Mathlib/NumberTheory/LegendreSymbol/AddCharacter.lean index a2077362c42e0..783acaa6f2fc5 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/AddCharacter.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/AddCharacter.lean @@ -95,16 +95,12 @@ def toMonoidHom : AddChar R R' → Multiplicative R →* R' := open Multiplicative --- Porting note: added. -@[coe] -def toFun (ψ : AddChar R R') (x : R) : R' := ψ.toMonoidHom (ofAdd x) - /-- Define coercion to a function so that it includes the move from `R` to `Multiplicative R`. After we have proved the API lemmas below, we don't need to worry about writing `ofAdd a` when we want to apply an additive character. -/ -instance hasCoeToFun : CoeFun (AddChar R R') fun _ => R → R' where - coe := toFun -#align add_char.has_coe_to_fun AddChar.hasCoeToFun +instance instFunLike : FunLike (AddChar R R') R fun _ ↦ R' := + inferInstanceAs (FunLike (Multiplicative R →* R') R fun _ ↦ R') +#noalign add_char.has_coe_to_fun theorem coe_to_fun_apply (ψ : AddChar R R') (a : R) : ψ a = ψ.toMonoidHom (ofAdd a) := rfl @@ -118,9 +114,8 @@ theorem mul_apply (ψ φ : AddChar R R') (a : R) : (ψ * φ) a = ψ a * φ a := @[simp] theorem one_apply (a : R) : (1 : AddChar R R') a = 1 := rfl -instance monoidHomClass : MonoidHomClass (AddChar R R') (Multiplicative R) R' := - MonoidHom.monoidHomClass -#align add_char.monoid_hom_class AddChar.monoidHomClass +-- this instance was a bad idea and conflicted with `instFunLike` above +#noalign add_char.monoid_hom_class -- Porting note(https://github.com/leanprover-community/mathlib4/issues/5229): added. @[ext] diff --git a/Mathlib/NumberTheory/LegendreSymbol/GaussSum.lean b/Mathlib/NumberTheory/LegendreSymbol/GaussSum.lean index 88581fa7e69e3..59cbd04d59fc9 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/GaussSum.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/GaussSum.lean @@ -296,9 +296,13 @@ theorem FiniteField.two_pow_card {F : Type*} [Fintype F] [Field F] (hF : ringCha let τ : FF := ψ₈char 1 have τ_spec : τ ^ 4 = -1 := by refine (sq_eq_one_iff.1 ?_).resolve_left ?_ - · rw [← pow_mul, ← map_nsmul_pow ψ₈char, AddChar.IsPrimitive.zmod_char_eq_one_iff 8 ψ₈.prim] + · rw [← pow_mul, ← map_nsmul_pow ψ₈char] + -- doesn't match syntactically for `rw` + refine (AddChar.IsPrimitive.zmod_char_eq_one_iff 8 ψ₈.prim _).2 ?_ decide - · rw [← map_nsmul_pow ψ₈char, AddChar.IsPrimitive.zmod_char_eq_one_iff 8 ψ₈.prim] + · rw [← map_nsmul_pow ψ₈char] + -- doesn't match syntactically for `rw` + refine (AddChar.IsPrimitive.zmod_char_eq_one_iff 8 ψ₈.prim _).not.2 ?_ decide -- we consider `χ₈` as a multiplicative character `ℤ/8ℤ → FF` diff --git a/Mathlib/NumberTheory/NumberField/Discriminant.lean b/Mathlib/NumberTheory/NumberField/Discriminant.lean index 5d31b17c6bc2a..70cdaf6747c82 100644 --- a/Mathlib/NumberTheory/NumberField/Discriminant.lean +++ b/Mathlib/NumberTheory/NumberField/Discriminant.lean @@ -48,7 +48,7 @@ theorem discr_eq_discr_of_algEquiv {L : Type*} [Field L] [NumberField L] (f : K refine Fintype.equivOfCardEq ?_ rw [← FiniteDimensional.finrank_eq_card_chooseBasisIndex, RingOfIntegers.rank, AlgHom.card] rw [← Rat.intCast_inj, coe_discr, Algebra.discr_eq_discr_of_algEquiv ℚ ℂ (integralBasis K) e f, - ← discr_eq_discr _ ((RingOfIntegers.basis K).map f₀)] + ← discr_eq_discr L ((RingOfIntegers.basis K).map f₀)] change _ = algebraMap ℤ ℚ _ rw [← Algebra.discr_localizationLocalization ℤ (nonZeroDivisors ℤ) L] congr diff --git a/Mathlib/NumberTheory/NumberField/Units.lean b/Mathlib/NumberTheory/NumberField/Units.lean index dba28f1050c07..404a8c0848175 100644 --- a/Mathlib/NumberTheory/NumberField/Units.lean +++ b/Mathlib/NumberTheory/NumberField/Units.lean @@ -466,7 +466,7 @@ open dirichletUnitTheorem FiniteDimensional Classical def rank : ℕ := Fintype.card (InfinitePlace K) - 1 instance instDiscrete_unitLattice : DiscreteTopology (unitLattice K) := by - refine discreteTopology_of_open_singleton_zero ?_ + refine discreteTopology_of_isOpen_singleton_zero ?_ refine isOpen_singleton_of_finite_mem_nhds 0 (s := Metric.closedBall 0 1) ?_ ?_ · exact Metric.closedBall_mem_nhds _ (by norm_num) · refine Set.Finite.of_finite_image ?_ (Set.injOn_of_injective Subtype.val_injective _) @@ -493,6 +493,7 @@ theorem unitLattice_rank : rw [← Units.finrank_eq_rank] exact Zlattice.rank ℝ (unitLattice_span_eq_top K) +set_option synthInstance.maxHeartbeats 27000 in /-- The linear equivalence between `unitLattice` and `(𝓞 K)ˣ ⧸ (torsion K)` as an additive `ℤ`-module. -/ def unitLatticeEquiv : (unitLattice K) ≃ₗ[ℤ] Additive ((𝓞 K)ˣ ⧸ (torsion K)) := by diff --git a/Mathlib/NumberTheory/Padics/PadicNumbers.lean b/Mathlib/NumberTheory/Padics/PadicNumbers.lean index 3d1f805187491..ae45a6bb5e9e4 100644 --- a/Mathlib/NumberTheory/Padics/PadicNumbers.lean +++ b/Mathlib/NumberTheory/Padics/PadicNumbers.lean @@ -3,6 +3,7 @@ Copyright (c) 2018 Robert Y. Lewis. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Robert Y. Lewis -/ +import Mathlib.RingTheory.Valuation.Basic import Mathlib.NumberTheory.Padics.PadicNorm import Mathlib.Analysis.Normed.Field.Basic @@ -601,7 +602,7 @@ theorem defn (f : PadicSeq p) {ε : ℚ} (hε : 0 < ε) : ∃ N, ∀ i ≥ N, padicNormE (Padic.mk f - f i : ℚ_[p]) < ε := by dsimp [padicNormE] change ∃ N, ∀ i ≥ N, (f - const _ (f i)).norm < ε - by_contra' h + by_contra! h cases' cauchy₂ f hε with N hN rcases h N with ⟨i, hi, hge⟩ have hne : ¬f - const (padicNorm p) (f i) ≈ 0 := fun h ↦ by @@ -710,7 +711,7 @@ theorem exi_rat_seq_conv_cauchy : IsCauSeq (padicNorm p) (limSeq f) := fun ε h exact mod_cast this · apply lt_of_le_of_lt · apply padicNormE.add_le - · rw [←add_thirds ε] + · rw [← add_thirds ε] apply _root_.add_lt_add · suffices padicNormE (limSeq f j - f j + (f j - f (max N N2)) : ℚ_[p]) < ε / 3 + ε / 3 by simpa only [sub_add_sub_cancel] @@ -992,7 +993,7 @@ theorem padicNormE_lim_le {f : CauSeq ℚ_[p] norm} {a : ℝ} (ha : 0 < a) (hf : -- Porting note: `Setoid.symm` cannot work out which `Setoid` to use, so instead swap the order -- now, I use a rewrite to swap it later obtain ⟨N, hN⟩ := (CauSeq.equiv_lim f) _ ha - rw [←sub_add_cancel f.lim (f N)] + rw [← sub_add_cancel f.lim (f N)] refine le_trans (padicNormE.nonarchimedean _ _) ?_ rw [norm_sub_rev] exact max_le (le_of_lt (hN _ le_rfl)) (hf _) diff --git a/Mathlib/NumberTheory/Padics/PadicVal.lean b/Mathlib/NumberTheory/Padics/PadicVal.lean index ba377c43d63e2..f953bec7900da 100644 --- a/Mathlib/NumberTheory/Padics/PadicVal.lean +++ b/Mathlib/NumberTheory/Padics/PadicVal.lean @@ -728,7 +728,7 @@ theorem sub_one_mul_padicValNat_choose_eq_sub_sum_digits' {k n : ℕ} [hp : Fact padicValNat.mul (factorial_ne_zero _) (factorial_ne_zero _), Nat.mul_add] simp only [sub_one_mul_padicValNat_factorial] rw [← Nat.sub_add_comm <| digit_sum_le p k, Nat.add_sub_cancel n k, ← Nat.add_sub_assoc <| - digit_sum_le p n, Nat.sub_sub (k + n), ← Nat.sub_right_comm, Nat.sub_sub, sub_add_eq, + digit_sum_le p n, Nat.sub_sub (k + n), ← Nat.sub_right_comm, Nat.sub_sub, sub_add_eq, add_comm, tsub_tsub_assoc (Nat.le_refl (k + n)) <| (add_comm k n) ▸ (Nat.add_le_add (digit_sum_le p n) (digit_sum_le p k)), Nat.sub_self (k + n), zero_add, add_comm] diff --git a/Mathlib/NumberTheory/Wilson.lean b/Mathlib/NumberTheory/Wilson.lean index 2405518e79d5d..719004dc51467 100644 --- a/Mathlib/NumberTheory/Wilson.lean +++ b/Mathlib/NumberTheory/Wilson.lean @@ -48,7 +48,7 @@ theorem wilsons_lemma : ((p - 1)! : ZMod p) = -1 := by -- simp_rw [← Units.coeHom_apply, ← (Units.coeHom (ZMod p)).map_prod, -- prod_univ_units_id_eq_neg_one, Units.coeHom_apply, Units.val_neg, Units.val_one] simp_rw [← Units.coeHom_apply] - rw [← (Units.coeHom (ZMod p)).map_prod] + rw [← map_prod (Units.coeHom (ZMod p))] simp_rw [prod_univ_units_id_eq_neg_one, Units.coeHom_apply, Units.val_neg, Units.val_one] have hp : 0 < p := (Fact.out (p := p.Prime)).pos symm diff --git a/Mathlib/Order/Birkhoff.lean b/Mathlib/Order/Birkhoff.lean index a8bdeedadec04..38a9573d1dd19 100644 --- a/Mathlib/Order/Birkhoff.lean +++ b/Mathlib/Order/Birkhoff.lean @@ -52,7 +52,7 @@ variable [SemilatticeInf α] {s : UpperSet α} {a : α} @[simp] lemma infIrred_Ici (a : α) : InfIrred (Ici a) := by refine ⟨fun h ↦ Ici_ne_top h.eq_top, fun s t hst ↦ ?_⟩ have := mem_Ici_iff.2 (le_refl a) - rw [←hst] at this + rw [← hst] at this exact this.imp (fun ha ↦ le_antisymm (le_Ici.2 ha) $ hst.ge.trans inf_le_left) fun ha ↦ le_antisymm (le_Ici.2 ha) $ hst.ge.trans inf_le_right @@ -74,7 +74,7 @@ variable [SemilatticeSup α] {s : LowerSet α} {a : α} @[simp] lemma supIrred_Iic (a : α) : SupIrred (Iic a) := by refine' ⟨fun h ↦ Iic_ne_bot h.eq_bot, fun s t hst ↦ _⟩ have := mem_Iic_iff.2 (le_refl a) - rw [←hst] at this + rw [← hst] at this exact this.imp (fun ha ↦ (le_sup_left.trans_eq hst).antisymm $ Iic_le.2 ha) fun ha ↦ (le_sup_right.trans_eq hst).antisymm $ Iic_le.2 ha diff --git a/Mathlib/Order/BooleanAlgebra.lean b/Mathlib/Order/BooleanAlgebra.lean index 37d38f5edf533..a2cbcd158a2ac 100644 --- a/Mathlib/Order/BooleanAlgebra.lean +++ b/Mathlib/Order/BooleanAlgebra.lean @@ -216,7 +216,7 @@ theorem disjoint_sdiff_self_right : Disjoint x (y \ x) := lemma le_sdiff : x ≤ y \ z ↔ x ≤ y ∧ Disjoint x z := ⟨fun h ↦ ⟨h.trans sdiff_le, disjoint_sdiff_self_left.mono_left h⟩, fun h ↦ - by rw [←h.2.sdiff_eq_left]; exact sdiff_le_sdiff_right h.1⟩ + by rw [← h.2.sdiff_eq_left]; exact sdiff_le_sdiff_right h.1⟩ #align le_sdiff le_sdiff @[simp] lemma sdiff_eq_left : x \ y = x ↔ Disjoint x y := diff --git a/Mathlib/Order/Bounds/Basic.lean b/Mathlib/Order/Bounds/Basic.lean index 9b0166eed10dc..1f4e082b79042 100644 --- a/Mathlib/Order/Bounds/Basic.lean +++ b/Mathlib/Order/Bounds/Basic.lean @@ -1577,7 +1577,7 @@ lemma bddBelow_pi {s : Set (∀ a, π a)} : lemma bddAbove_range_pi {F : ι → ∀ a, π a} : BddAbove (range F) ↔ ∀ a, BddAbove (range fun i ↦ F i a) := by - simp only [bddAbove_pi, ←range_comp] + simp only [bddAbove_pi, ← range_comp] rfl lemma bddBelow_range_pi {F : ι → ∀ a, π a} : diff --git a/Mathlib/Order/Category/NonemptyFinLinOrd.lean b/Mathlib/Order/Category/NonemptyFinLinOrd.lean index 43a034041790e..56cfc23252b87 100644 --- a/Mathlib/Order/Category/NonemptyFinLinOrd.lean +++ b/Mathlib/Order/Category/NonemptyFinLinOrd.lean @@ -170,7 +170,7 @@ theorem epi_iff_surjective {A B : NonemptyFinLinOrd.{u}} (f : A ⟶ B) : constructor · intro dsimp only [Function.Surjective] - by_contra' hf' + by_contra! hf' rcases hf' with ⟨m, hm⟩ let Y := NonemptyFinLinOrd.of (ULift (Fin 2)) let p₁ : B ⟶ Y := diff --git a/Mathlib/Order/Closure.lean b/Mathlib/Order/Closure.lean index 281b32c20f8c6..2b9f73336ed62 100644 --- a/Mathlib/Order/Closure.lean +++ b/Mathlib/Order/Closure.lean @@ -217,7 +217,7 @@ theorem eq_mk₃_closed (c : ClosureOperator α) : @[simp] theorem mem_mk₃_closed {f : α → α} {p : α → Prop} {hf hfp hmin} {x : α} : x ∈ (mk₃ f p hf hfp hmin).closed ↔ p x := by refine' ⟨λ hx ↦ _, λ hx ↦ (hmin le_rfl hx).antisymm (hf _)⟩ - rw [←closure_eq_self_of_mem_closed _ hx] + rw [← closure_eq_self_of_mem_closed _ hx] exact hfp _ #align closure_operator.mem_mk₃_closed ClosureOperator.mem_mk₃_closedₓ diff --git a/Mathlib/Order/CompleteLattice.lean b/Mathlib/Order/CompleteLattice.lean index 69104de780290..a469d4826158c 100644 --- a/Mathlib/Order/CompleteLattice.lean +++ b/Mathlib/Order/CompleteLattice.lean @@ -1946,7 +1946,7 @@ theorem snd_iInf [InfSet α] [InfSet β] (f : ι → α × β) : (iInf f).snd = #align prod.snd_infi Prod.snd_iInf theorem swap_iInf [InfSet α] [InfSet β] (f : ι → α × β) : (iInf f).swap = ⨅ i, (f i).swap := by - simp_rw [iInf, swap_sInf, ←range_comp, Function.comp] -- Porting note: need to unfold `∘` + simp_rw [iInf, swap_sInf, ← range_comp, Function.comp] -- Porting note: need to unfold `∘` #align prod.swap_infi Prod.swap_iInf theorem iInf_mk [InfSet α] [InfSet β] (f : ι → α) (g : ι → β) : @@ -1963,7 +1963,7 @@ theorem snd_iSup [SupSet α] [SupSet β] (f : ι → α × β) : (iSup f).snd = #align prod.snd_supr Prod.snd_iSup theorem swap_iSup [SupSet α] [SupSet β] (f : ι → α × β) : (iSup f).swap = ⨆ i, (f i).swap := by - simp_rw [iSup, swap_sSup, ←range_comp, Function.comp] -- Porting note: need to unfold `∘` + simp_rw [iSup, swap_sSup, ← range_comp, Function.comp] -- Porting note: need to unfold `∘` #align prod.swap_supr Prod.swap_iSup theorem iSup_mk [SupSet α] [SupSet β] (f : ι → α) (g : ι → β) : diff --git a/Mathlib/Order/CompleteLatticeIntervals.lean b/Mathlib/Order/CompleteLatticeIntervals.lean index 5ded663285a01..77ed49526ef6f 100644 --- a/Mathlib/Order/CompleteLatticeIntervals.lean +++ b/Mathlib/Order/CompleteLatticeIntervals.lean @@ -187,7 +187,7 @@ noncomputable def Set.Icc.completeLattice [ConditionallyCompleteLattice α] {a b : α} (h : a ≤ b) : CompleteLattice (Set.Icc a b) where __ := Set.Icc.boundedOrder h sSup S := if hS : S = ∅ then ⟨a, le_rfl, h⟩ else ⟨sSup ((↑) '' S), by - rw [←Set.not_nonempty_iff_eq_empty, not_not] at hS + rw [← Set.not_nonempty_iff_eq_empty, not_not] at hS refine' ⟨_, csSup_le (hS.image (↑)) (fun _ ⟨c, _, hc⟩ ↦ hc ▸ c.2.2)⟩ obtain ⟨c, hc⟩ := hS exact c.2.1.trans (le_csSup ⟨b, fun _ ⟨d, _, hd⟩ ↦ hd ▸ d.2.2⟩ ⟨c, hc, rfl⟩)⟩ @@ -201,7 +201,7 @@ noncomputable def Set.Icc.completeLattice [ConditionallyCompleteLattice α] · exact csSup_le ((Set.nonempty_iff_ne_empty.mpr hS).image (↑)) (fun _ ⟨d, h, hd⟩ ↦ hd ▸ hc d h) sInf S := if hS : S = ∅ then ⟨b, h, le_rfl⟩ else ⟨sInf ((↑) '' S), by - rw [←Set.not_nonempty_iff_eq_empty, not_not] at hS + rw [← Set.not_nonempty_iff_eq_empty, not_not] at hS refine' ⟨le_csInf (hS.image (↑)) (fun _ ⟨c, _, hc⟩ ↦ hc ▸ c.2.1), _⟩ obtain ⟨c, hc⟩ := hS exact le_trans (csInf_le ⟨a, fun _ ⟨d, _, hd⟩ ↦ hd ▸ d.2.1⟩ ⟨c, hc, rfl⟩) c.2.2⟩ diff --git a/Mathlib/Order/Extension/Linear.lean b/Mathlib/Order/Extension/Linear.lean index c9ef939ccc920..4fdb7b279deb9 100644 --- a/Mathlib/Order/Extension/Linear.lean +++ b/Mathlib/Order/Extension/Linear.lean @@ -53,7 +53,7 @@ theorem extend_partialOrder {α : Type u} (r : α → α → Prop) [IsPartialOrd haveI : IsPartialOrder α s := hs₁ refine ⟨s, { total := ?_, refl := hs₁.refl, trans := hs₁.trans, antisymm := hs₁.antisymm } , rs⟩ intro x y - by_contra' h + by_contra! h let s' x' y' := s x' y' ∨ s x' x ∧ s y y' rw [← hs₂ s' _ fun _ _ ↦ Or.inl] at h · apply h.1 (Or.inr ⟨refl _, refl _⟩) diff --git a/Mathlib/Order/Filter/Archimedean.lean b/Mathlib/Order/Filter/Archimedean.lean index ccfa4a1949c50..7039000f2e27a 100644 --- a/Mathlib/Order/Filter/Archimedean.lean +++ b/Mathlib/Order/Filter/Archimedean.lean @@ -74,11 +74,11 @@ theorem tendsto_int_cast_atTop_atTop [StrictOrderedRing R] [Archimedean R] : theorem Filter.Eventually.int_cast_atTop [StrictOrderedRing R] [Archimedean R] {p : R → Prop} (h : ∀ᶠ (x:R) in atTop, p x) : ∀ᶠ (n:ℤ) in atTop, p n := by - rw [←Int.comap_cast_atTop (R := R)]; exact h.comap _ + rw [← Int.comap_cast_atTop (R := R)]; exact h.comap _ theorem Filter.Eventually.int_cast_atBot [StrictOrderedRing R] [Archimedean R] {p : R → Prop} (h : ∀ᶠ (x:R) in atBot, p x) : ∀ᶠ (n:ℤ) in atBot, p n := by - rw [←Int.comap_cast_atBot (R := R)]; exact h.comap _ + rw [← Int.comap_cast_atBot (R := R)]; exact h.comap _ @[simp] theorem Rat.comap_cast_atTop [LinearOrderedField R] [Archimedean R] : @@ -106,11 +106,11 @@ theorem tendsto_rat_cast_atBot_iff [LinearOrderedField R] [Archimedean R] {f : theorem Filter.Eventually.rat_cast_atTop [LinearOrderedField R] [Archimedean R] {p : R → Prop} (h : ∀ᶠ (x:R) in atTop, p x) : ∀ᶠ (n:ℚ) in atTop, p n := by - rw [←Rat.comap_cast_atTop (R := R)]; exact h.comap _ + rw [← Rat.comap_cast_atTop (R := R)]; exact h.comap _ theorem Filter.Eventually.rat_cast_atBot [LinearOrderedField R] [Archimedean R] {p : R → Prop} (h : ∀ᶠ (x:R) in atBot, p x) : ∀ᶠ (n:ℚ) in atBot, p n := by - rw [←Rat.comap_cast_atBot (R := R)]; exact h.comap _ + rw [← Rat.comap_cast_atBot (R := R)]; exact h.comap _ -- porting note: new lemma theorem atTop_hasAntitoneBasis_of_archimedean [OrderedSemiring R] [Archimedean R] : diff --git a/Mathlib/Order/Filter/AtTopBot.lean b/Mathlib/Order/Filter/AtTopBot.lean index f831f2f24efc2..ecb105a1add78 100644 --- a/Mathlib/Order/Filter/AtTopBot.lean +++ b/Mathlib/Order/Filter/AtTopBot.lean @@ -264,12 +264,12 @@ theorem eventually_forall_le_atBot [Preorder α] {p : α → Prop} : theorem Tendsto.eventually_forall_ge_atTop {α β : Type*} [Preorder β] {l : Filter α} {p : β → Prop} {f : α → β} (hf : Tendsto f l atTop) (h_evtl : ∀ᶠ x in atTop, p x) : ∀ᶠ x in l, ∀ y, f x ≤ y → p y := by - rw [←Filter.eventually_forall_ge_atTop] at h_evtl; exact (h_evtl.comap f).filter_mono hf.le_comap + rw [← Filter.eventually_forall_ge_atTop] at h_evtl; exact (h_evtl.comap f).filter_mono hf.le_comap theorem Tendsto.eventually_forall_le_atBot {α β : Type*} [Preorder β] {l : Filter α} {p : β → Prop} {f : α → β} (hf : Tendsto f l atBot) (h_evtl : ∀ᶠ x in atBot, p x) : ∀ᶠ x in l, ∀ y, y ≤ f x → p y := by - rw [←Filter.eventually_forall_le_atBot] at h_evtl; exact (h_evtl.comap f).filter_mono hf.le_comap + rw [← Filter.eventually_forall_le_atBot] at h_evtl; exact (h_evtl.comap f).filter_mono hf.le_comap theorem atTop_basis_Ioi [Nonempty α] [SemilatticeSup α] [NoMaxOrder α] : (@atTop α _).HasBasis (fun _ => True) Ioi := diff --git a/Mathlib/Order/Filter/Basic.lean b/Mathlib/Order/Filter/Basic.lean index bb04c7a7000fa..16d2a914c08d2 100644 --- a/Mathlib/Order/Filter/Basic.lean +++ b/Mathlib/Order/Filter/Basic.lean @@ -2930,7 +2930,7 @@ gi_principal_ker.gc.u_iInf gi_principal_ker.gc.u_sInf @[simp] lemma ker_principal (s : Set α) : ker (𝓟 s) = s := gi_principal_ker.u_l_eq _ -@[simp] lemma ker_pure (a : α) : ker (pure a) = {a} := by rw [←principal_singleton, ker_principal] +@[simp] lemma ker_pure (a : α) : ker (pure a) = {a} := by rw [← principal_singleton, ker_principal] @[simp] lemma ker_comap (m : α → β) (f : Filter β) : ker (comap m f) = m ⁻¹' ker f := by ext a diff --git a/Mathlib/Order/Filter/Cofinite.lean b/Mathlib/Order/Filter/Cofinite.lean index f0ae40dc18a56..5654ad5b45923 100644 --- a/Mathlib/Order/Filter/Cofinite.lean +++ b/Mathlib/Order/Filter/Cofinite.lean @@ -146,7 +146,7 @@ theorem countable_compl_ker [l.IsCountablyGenerated] (h : cofinite ≤ l) : Set. then for all but countably many elements, `f x ∈ l.ker`. -/ theorem Tendsto.countable_compl_preimage_ker {f : α → β} {l : Filter β} [l.IsCountablyGenerated] (h : Tendsto f cofinite l) : - Set.Countable (f ⁻¹' l.ker)ᶜ := by rw [←ker_comap]; exact countable_compl_ker h.le_comap + Set.Countable (f ⁻¹' l.ker)ᶜ := by rw [← ker_comap]; exact countable_compl_ker h.le_comap end Filter diff --git a/Mathlib/Order/Filter/Germ.lean b/Mathlib/Order/Filter/Germ.lean index db658aaca5937..e16cd98d2ab92 100644 --- a/Mathlib/Order/Filter/Germ.lean +++ b/Mathlib/Order/Filter/Germ.lean @@ -457,14 +457,16 @@ theorem coe_nat [NatCast M] (n : ℕ) : ((fun _ ↦ n : α → M) : Germ l M) = @[simp, norm_cast] theorem const_nat [NatCast M] (n : ℕ) : ((n : M) : Germ l M) = n := rfl +-- See note [no_index around OfNat.ofNat] @[simp, norm_cast] theorem coe_ofNat [NatCast M] (n : ℕ) [n.AtLeastTwo] : - ((OfNat.ofNat n : α → M) : Germ l M) = OfNat.ofNat n := + ((no_index (OfNat.ofNat n : α → M)) : Germ l M) = OfNat.ofNat n := rfl +-- See note [no_index around OfNat.ofNat] @[simp, norm_cast] theorem const_ofNat [NatCast M] (n : ℕ) [n.AtLeastTwo] : - ((OfNat.ofNat n : M) : Germ l M) = OfNat.ofNat n := + ((no_index (OfNat.ofNat n : M)) : Germ l M) = OfNat.ofNat n := rfl instance intCast [IntCast M] : IntCast (Germ l M) where diff --git a/Mathlib/Order/Filter/Partial.lean b/Mathlib/Order/Filter/Partial.lean index 201be27938299..1cc8650e67962 100644 --- a/Mathlib/Order/Filter/Partial.lean +++ b/Mathlib/Order/Filter/Partial.lean @@ -62,7 +62,7 @@ def rmap (r : Rel α β) (l : Filter α) : Filter β where inter_sets hs ht := by simp only [Set.mem_setOf_eq] convert inter_mem hs ht - rw [←Rel.core_inter] + rw [← Rel.core_inter] #align filter.rmap Filter.rmap theorem rmap_sets (r : Rel α β) (l : Filter α) : (l.rmap r).sets = r.core ⁻¹' l.sets := @@ -130,7 +130,7 @@ theorem rcomap_compose (r : Rel α β) (s : Rel β γ) : rcomap r ∘ rcomap s = theorem rtendsto_iff_le_rcomap (r : Rel α β) (l₁ : Filter α) (l₂ : Filter β) : RTendsto r l₁ l₂ ↔ l₁ ≤ l₂.rcomap r := by rw [rtendsto_def] - simp_rw [←l₂.mem_sets] + simp_rw [← l₂.mem_sets] simp [Filter.le_def, rcomap, Rel.mem_image]; constructor · exact fun h s t tl₂ => mem_of_superset (h t tl₂) · exact fun h t tl₂ => h _ t tl₂ Set.Subset.rfl diff --git a/Mathlib/Order/Hom/Basic.lean b/Mathlib/Order/Hom/Basic.lean index 13d8af2094d55..06eee5893eafb 100644 --- a/Mathlib/Order/Hom/Basic.lean +++ b/Mathlib/Order/Hom/Basic.lean @@ -886,8 +886,11 @@ theorem symm_symm (e : α ≃o β) : e.symm.symm = e := by rfl #align order_iso.symm_symm OrderIso.symm_symm -theorem symm_injective : Function.Injective (symm : α ≃o β → β ≃o α) := fun e e' h => by - rw [← e.symm_symm, h, e'.symm_symm] +theorem symm_bijective : Function.Bijective (OrderIso.symm : (α ≃o β) → β ≃o α) := + Function.bijective_iff_has_inverse.mpr ⟨_, symm_symm, symm_symm⟩ + +theorem symm_injective : Function.Injective (symm : α ≃o β → β ≃o α) := + symm_bijective.injective #align order_iso.symm_injective OrderIso.symm_injective @[simp] diff --git a/Mathlib/Order/Lattice.lean b/Mathlib/Order/Lattice.lean index 223bb8e80ab7c..f550108409de8 100644 --- a/Mathlib/Order/Lattice.lean +++ b/Mathlib/Order/Lattice.lean @@ -57,20 +57,7 @@ universe u v w variable {α : Type u} {β : Type v} --- TODO: move this eventually, if we decide to use them --- Porting note: no ematch attribute ---attribute [ematch] le_trans lt_of_le_of_lt lt_of_lt_of_le lt_trans - -section - --- TODO: this seems crazy, but it also seems to work reasonably well --- Porting note: no ematch attribute ---@[ematch] -theorem le_antisymm' [PartialOrder α] : ∀ {a b : α}, a ≤ b → b ≤ a → a = b := - @le_antisymm _ _ -#align le_antisymm' le_antisymm' - -end +#align le_antisymm' le_antisymm /-! ### Join-semilattices @@ -684,7 +671,7 @@ theorem inf_le_sup : a ⊓ b ≤ a ⊔ b := theorem sup_le_inf : a ⊔ b ≤ a ⊓ b ↔ a = b := by simp [le_antisymm_iff, and_comm] #align sup_le_inf sup_le_inf -@[simp] lemma inf_eq_sup : a ⊓ b = a ⊔ b ↔ a = b := by rw [←inf_le_sup.ge_iff_eq, sup_le_inf] +@[simp] lemma inf_eq_sup : a ⊓ b = a ⊔ b ↔ a = b := by rw [← inf_le_sup.ge_iff_eq, sup_le_inf] #align inf_eq_sup inf_eq_sup @[simp] lemma sup_eq_inf : a ⊔ b = a ⊓ b ↔ a = b := eq_comm.trans inf_eq_sup #align sup_eq_inf sup_eq_inf diff --git a/Mathlib/Order/LiminfLimsup.lean b/Mathlib/Order/LiminfLimsup.lean index 9afa703f068db..0fdd22538d53e 100644 --- a/Mathlib/Order/LiminfLimsup.lean +++ b/Mathlib/Order/LiminfLimsup.lean @@ -1164,7 +1164,7 @@ lemma mem_liminf_iff_eventually_mem : (a ∈ liminf s 𝓕) ↔ (∀ᶠ i in using ⟨fun ⟨S, hS, hS'⟩ ↦ mem_of_superset hS (by tauto), fun h ↦ ⟨{i | a ∈ s i}, h, by tauto⟩⟩ lemma mem_limsup_iff_frequently_mem : (a ∈ limsup s 𝓕) ↔ (∃ᶠ i in 𝓕, a ∈ s i) := by - simp only [Filter.Frequently, iff_not_comm, ←mem_compl_iff, limsup_compl, comp_apply, + simp only [Filter.Frequently, iff_not_comm, ← mem_compl_iff, limsup_compl, comp_apply, mem_liminf_iff_eventually_mem] theorem cofinite.blimsup_set_eq : diff --git a/Mathlib/Order/Minimal.lean b/Mathlib/Order/Minimal.lean index d037a1cc0c026..0f160ecab2bfc 100644 --- a/Mathlib/Order/Minimal.lean +++ b/Mathlib/Order/Minimal.lean @@ -339,8 +339,8 @@ theorem inter_minimals_preimage_inter_eq_of_rel_iff_rel_on theorem inter_preimage_minimals_eq_of_rel_iff_rel_on_of_subset (hf : ∀ ⦃a a'⦄, a ∈ x → a' ∈ x → (r a a' ↔ s (f a) (f a'))) (hy : y ⊆ f '' x) : x ∩ f ⁻¹' (minimals s y) = minimals r (x ∩ f ⁻¹' y) := by - rw [←inter_eq_self_of_subset_right hy, inter_minimals_preimage_inter_eq_of_rel_iff_rel_on hf, - preimage_inter, ←inter_assoc, inter_eq_self_of_subset_left (subset_preimage_image f x)] + rw [← inter_eq_self_of_subset_right hy, inter_minimals_preimage_inter_eq_of_rel_iff_rel_on hf, + preimage_inter, ← inter_assoc, inter_eq_self_of_subset_left (subset_preimage_image f x)] theorem RelEmbedding.inter_preimage_minimals_eq (f : r ↪r s) (x : Set α) (y : Set β) : x ∩ f⁻¹' (minimals s ((f '' x) ∩ y)) = minimals r (x ∩ f ⁻¹' y) := diff --git a/Mathlib/Order/Monotone/Basic.lean b/Mathlib/Order/Monotone/Basic.lean index d2a8a70182ae6..4d08036fad4d8 100644 --- a/Mathlib/Order/Monotone/Basic.lean +++ b/Mathlib/Order/Monotone/Basic.lean @@ -959,7 +959,7 @@ downright. -/ lemma not_monotone_not_antitone_iff_exists_lt_lt : ¬ Monotone f ∧ ¬ Antitone f ↔ ∃ a b c, a < b ∧ b < c ∧ (f a < f b ∧ f c < f b ∨ f b < f a ∧ f b < f c) := by - simp_rw [not_monotone_not_antitone_iff_exists_le_le, ←and_assoc] + simp_rw [not_monotone_not_antitone_iff_exists_le_le, ← and_assoc] refine' exists₃_congr (fun a b c ↦ and_congr_left $ fun h ↦ (Ne.le_iff_lt _).and $ Ne.le_iff_lt _) <;> (rintro rfl; simp at h) diff --git a/Mathlib/Order/OmegaCompletePartialOrder.lean b/Mathlib/Order/OmegaCompletePartialOrder.lean index 39da9f03f9bb1..bc6f5cd05c86c 100644 --- a/Mathlib/Order/OmegaCompletePartialOrder.lean +++ b/Mathlib/Order/OmegaCompletePartialOrder.lean @@ -566,7 +566,7 @@ variable {α β : Type*} [OmegaCompletePartialOrder α] [CompleteLinearOrder β] theorem inf_continuous (f g : α →o β) (hf : Continuous f) (hg : Continuous g) : Continuous (f ⊓ g) := by refine' fun c => eq_of_forall_ge_iff fun z => _ - simp only [inf_le_iff, hf c, hg c, ωSup_le_iff, ←forall_or_left, ←forall_or_right, + simp only [inf_le_iff, hf c, hg c, ωSup_le_iff, ← forall_or_left, ← forall_or_right, Chain.map_coe, OrderHom.coe_inf, ge_iff_le, Pi.inf_apply, Function.comp] exact ⟨λ h _ => h _ _, λ h i j => (h (max j i)).imp (le_trans $ f.mono $ c.mono $ le_max_left _ _) (le_trans $ g.mono $ c.mono $ le_max_right _ _)⟩ diff --git a/Mathlib/Order/PartialSups.lean b/Mathlib/Order/PartialSups.lean index e04b1ff22b5c5..028cc1cfa2793 100644 --- a/Mathlib/Order/PartialSups.lean +++ b/Mathlib/Order/PartialSups.lean @@ -122,6 +122,10 @@ theorem partialSups_eq_sup'_range (f : ℕ → α) (n : ℕ) : eq_of_forall_ge_iff fun _ ↦ by simp [Nat.lt_succ_iff] #align partial_sups_eq_sup'_range partialSups_eq_sup'_range +lemma partialSups_apply {ι : Type*} {π : ι → Type*} [(i : ι) → SemilatticeSup (π i)] + (f : ℕ → (i : ι) → π i) (n : ℕ) (i : ι) : partialSups f n i = partialSups (f · i) n := by + simp only [partialSups_eq_sup'_range, Finset.sup'_apply] + end SemilatticeSup theorem partialSups_eq_sup_range [SemilatticeSup α] [OrderBot α] (f : ℕ → α) (n : ℕ) : diff --git a/Mathlib/Order/RelClasses.lean b/Mathlib/Order/RelClasses.lean index 2e8bf76bb5049..6c325d05603e8 100644 --- a/Mathlib/Order/RelClasses.lean +++ b/Mathlib/Order/RelClasses.lean @@ -620,7 +620,7 @@ variable [HasSubset α] {a b c : α} lemma subset_of_eq_of_subset (hab : a = b) (hbc : b ⊆ c) : a ⊆ c := by rwa [hab] #align subset_of_eq_of_subset subset_of_eq_of_subset -lemma subset_of_subset_of_eq (hab : a ⊆ b) (hbc : b = c) : a ⊆ c := by rwa [←hbc] +lemma subset_of_subset_of_eq (hab : a ⊆ b) (hbc : b = c) : a ⊆ c := by rwa [← hbc] #align subset_of_subset_of_eq subset_of_subset_of_eq @[refl] @@ -689,7 +689,7 @@ variable [HasSSubset α] {a b c : α} lemma ssubset_of_eq_of_ssubset (hab : a = b) (hbc : b ⊂ c) : a ⊂ c := by rwa [hab] #align ssubset_of_eq_of_ssubset ssubset_of_eq_of_ssubset -lemma ssubset_of_ssubset_of_eq (hab : a ⊂ b) (hbc : b = c) : a ⊂ c := by rwa [←hbc] +lemma ssubset_of_ssubset_of_eq (hab : a ⊂ b) (hbc : b = c) : a ⊂ c := by rwa [← hbc] #align ssubset_of_ssubset_of_eq ssubset_of_ssubset_of_eq lemma ssubset_irrefl [IsIrrefl α (· ⊂ ·)] (a : α) : ¬a ⊂ a := irrefl _ diff --git a/Mathlib/Order/Sublattice.lean b/Mathlib/Order/Sublattice.lean index 7e5524b0b3b0e..e0b20890d652d 100644 --- a/Mathlib/Order/Sublattice.lean +++ b/Mathlib/Order/Sublattice.lean @@ -165,16 +165,16 @@ def topEquiv : (⊤ : Sublattice α) ≃o α where @[simp, norm_cast] lemma coe_iInf (f : ι → Sublattice α) : ⨅ i, f i = ⋂ i, (f i : Set α) := by simp [iInf] -@[simp, norm_cast] lemma coe_eq_univ : L = (univ : Set α) ↔ L = ⊤ := by rw [←coe_top, coe_inj] -@[simp, norm_cast] lemma coe_eq_empty : L = (∅ : Set α) ↔ L = ⊥ := by rw [←coe_bot, coe_inj] +@[simp, norm_cast] lemma coe_eq_univ : L = (univ : Set α) ↔ L = ⊤ := by rw [← coe_top, coe_inj] +@[simp, norm_cast] lemma coe_eq_empty : L = (∅ : Set α) ↔ L = ⊥ := by rw [← coe_bot, coe_inj] @[simp] lemma not_mem_bot (a : α) : a ∉ (⊥ : Sublattice α) := id @[simp] lemma mem_top (a : α) : a ∈ (⊤ : Sublattice α) := mem_univ _ @[simp] lemma mem_inf : a ∈ L ⊓ M ↔ a ∈ L ∧ a ∈ M := Iff.rfl @[simp] lemma mem_sInf {S : Set (Sublattice α)} : a ∈ sInf S ↔ ∀ L ∈ S, a ∈ L := by - rw [←SetLike.mem_coe]; simp + rw [← SetLike.mem_coe]; simp @[simp] lemma mem_iInf {f : ι → Sublattice α} : a ∈ ⨅ i, f i ↔ ∀ i, a ∈ f i := by - rw [←SetLike.mem_coe]; simp + rw [← SetLike.mem_coe]; simp /-- Sublattices of a lattice form a complete lattice. -/ instance instCompleteLattice : CompleteLattice (Sublattice α) where @@ -246,7 +246,7 @@ lemma comap_equiv_eq_map_symm (f : β ≃o α) (L : Sublattice α) : lemma map_symm_eq_iff_eq_map {M : Sublattice β} {e : β ≃o α} : L.map ↑e.symm = M ↔ L = M.map ↑e := by - simp_rw [←coe_inj]; exact (Equiv.eq_image_iff_symm_image_eq _ _ _).symm + simp_rw [← coe_inj]; exact (Equiv.eq_image_iff_symm_image_eq _ _ _).symm lemma map_le_iff_le_comap {f : LatticeHom α β} {M : Sublattice β} : L.map f ≤ M ↔ L ≤ M.comap f := image_subset_iff diff --git a/Mathlib/Order/SuccPred/LinearLocallyFinite.lean b/Mathlib/Order/SuccPred/LinearLocallyFinite.lean index da8b1ec6b6b4d..1a7f0ced832d4 100644 --- a/Mathlib/Order/SuccPred/LinearLocallyFinite.lean +++ b/Mathlib/Order/SuccPred/LinearLocallyFinite.lean @@ -134,8 +134,7 @@ instance (priority := 100) LinearLocallyFiniteOrder.isSuccArchimedean [LocallyFi swap · refine' ⟨0, _⟩ simpa only [Function.iterate_zero, id.def] using hij - by_contra h - push_neg at h + by_contra! h have h_lt : ∀ n, succ^[n] i < j := by intro n induction' n with n hn diff --git a/Mathlib/Order/SupClosed.lean b/Mathlib/Order/SupClosed.lean index 4a32fc9cf4b34..c485fbbdd6b49 100644 --- a/Mathlib/Order/SupClosed.lean +++ b/Mathlib/Order/SupClosed.lean @@ -64,7 +64,7 @@ lemma SupClosed.preimage [SupHomClass F β α] (hs : SupClosed s) (f : F) : SupC lemma SupClosed.image [SupHomClass F α β] (hs : SupClosed s) (f : F) : SupClosed (f '' s) := by rintro _ ⟨a, ha, rfl⟩ _ ⟨b, hb, rfl⟩ - rw [←map_sup] + rw [← map_sup] exact Set.mem_image_of_mem _ $ hs ha hb lemma supClosed_range [SupHomClass F α β] (f : F) : SupClosed (Set.range f) := by @@ -128,7 +128,7 @@ lemma InfClosed.preimage [InfHomClass F β α] (hs : InfClosed s) (f : F) : InfC lemma InfClosed.image [InfHomClass F α β] (hs : InfClosed s) (f : F) : InfClosed (f '' s) := by rintro _ ⟨a, ha, rfl⟩ _ ⟨b, hb, rfl⟩ - rw [←map_inf] + rw [← map_inf] exact Set.mem_image_of_mem _ $ hs ha hb lemma infClosed_range [InfHomClass F α β] (f : F) : InfClosed (Set.range f) := by @@ -446,7 +446,7 @@ protected lemma InfClosed.supClosure (hs : InfClosed s) : InfClosed (supClosure ⟨supClosed_supClosure.infClosure, infClosed_infClosure⟩ lemma Set.Finite.latticeClosure (hs : s.Finite) : (latticeClosure s).Finite := by - rw [←supClosure_infClosure]; exact hs.infClosure.supClosure + rw [← supClosure_infClosure]; exact hs.infClosure.supClosure end DistribLattice diff --git a/Mathlib/Order/SupIndep.lean b/Mathlib/Order/SupIndep.lean index d6f66d6b6396e..54c5af38d5118 100644 --- a/Mathlib/Order/SupIndep.lean +++ b/Mathlib/Order/SupIndep.lean @@ -422,7 +422,7 @@ theorem independent_ne_bot_iff_independent : theorem Independent.injOn (ht : Independent t) : InjOn t {i | t i ≠ ⊥} := by rintro i _ j (hj : t j ≠ ⊥) h - by_contra' contra + by_contra! contra apply hj suffices t j ≤ ⨆ (k) (_ : k ≠ i), t k by replace ht := (ht i).mono_right this diff --git a/Mathlib/Order/UpperLower/Basic.lean b/Mathlib/Order/UpperLower/Basic.lean index 200d6d793cf47..3fcd97344ebc5 100644 --- a/Mathlib/Order/UpperLower/Basic.lean +++ b/Mathlib/Order/UpperLower/Basic.lean @@ -416,7 +416,7 @@ section LinearOrder variable [LinearOrder α] {s t : Set α} theorem IsUpperSet.total (hs : IsUpperSet s) (ht : IsUpperSet t) : s ⊆ t ∨ t ⊆ s := by - by_contra' h + by_contra! h simp_rw [Set.not_subset] at h obtain ⟨⟨a, has, hat⟩, b, hbt, hbs⟩ := h obtain hab | hba := le_total a b @@ -1660,9 +1660,9 @@ lemma sdiff_le_left : s.sdiff t ≤ s := diff_subset _ _ lemma erase_le : s.erase a ≤ s := diff_subset _ _ @[simp] protected lemma sdiff_eq_left : s.sdiff t = s ↔ Disjoint ↑s t := by - simp [←SetLike.coe_set_eq] + simp [← SetLike.coe_set_eq] -@[simp] lemma erase_eq : s.erase a = s ↔ a ∉ s := by rw [←sdiff_singleton]; simp [-sdiff_singleton] +@[simp] lemma erase_eq : s.erase a = s ↔ a ∉ s := by rw [← sdiff_singleton]; simp [-sdiff_singleton] @[simp] lemma sdiff_lt_left : s.sdiff t < s ↔ ¬ Disjoint ↑s t := sdiff_le_left.lt_iff_ne.trans LowerSet.sdiff_eq_left.not @@ -1688,7 +1688,7 @@ lemma lowerClosure_sup_sdiff (hts : t ⊆ s) (hst : ∀ b ∈ s, ∀ c ∈ t, c lowerClosure t ⊔ s.sdiff t = s := by rw [sup_comm, sdiff_sup_lowerClosure hts hst] lemma erase_sup_Iic (ha : a ∈ s) (has : ∀ b ∈ s, a ≤ b → b = a) : s.erase a ⊔ Iic a = s := by - rw [←lowerClosure_singleton, ←sdiff_singleton, sdiff_sup_lowerClosure] <;> simpa + rw [← lowerClosure_singleton, ← sdiff_singleton, sdiff_sup_lowerClosure] <;> simpa lemma Iic_sup_erase (ha : a ∈ s) (has : ∀ b ∈ s, a ≤ b → b = a) : Iic a ⊔ s.erase a = s := by rw [sup_comm, erase_sup_Iic ha has] @@ -1721,9 +1721,9 @@ lemma le_sdiff_left : s ≤ s.sdiff t := diff_subset _ _ lemma le_erase : s ≤ s.erase a := diff_subset _ _ @[simp] protected lemma sdiff_eq_left : s.sdiff t = s ↔ Disjoint ↑s t := by - simp [←SetLike.coe_set_eq] + simp [← SetLike.coe_set_eq] -@[simp] lemma erase_eq : s.erase a = s ↔ a ∉ s := by rw [←sdiff_singleton]; simp [-sdiff_singleton] +@[simp] lemma erase_eq : s.erase a = s ↔ a ∉ s := by rw [← sdiff_singleton]; simp [-sdiff_singleton] @[simp] lemma lt_sdiff_left : s < s.sdiff t ↔ ¬ Disjoint ↑s t := le_sdiff_left.gt_iff_ne.trans UpperSet.sdiff_eq_left.not @@ -1749,7 +1749,7 @@ lemma upperClosure_inf_sdiff (hts : t ⊆ s) (hst : ∀ b ∈ s, ∀ c ∈ t, b upperClosure t ⊓ s.sdiff t = s := by rw [inf_comm, sdiff_inf_upperClosure hts hst] lemma erase_inf_Ici (ha : a ∈ s) (has : ∀ b ∈ s, b ≤ a → b = a) : s.erase a ⊓ Ici a = s := by - rw [←upperClosure_singleton, ←sdiff_singleton, sdiff_inf_upperClosure] <;> simpa + rw [← upperClosure_singleton, ← sdiff_singleton, sdiff_inf_upperClosure] <;> simpa lemma Ici_inf_erase (ha : a ∈ s) (has : ∀ b ∈ s, b ≤ a → b = a) : Ici a ⊓ s.erase a = s := by rw [inf_comm, erase_inf_Ici ha has] diff --git a/Mathlib/Order/WellFounded.lean b/Mathlib/Order/WellFounded.lean index e07a25edd66f3..f92e43ee2cf9a 100644 --- a/Mathlib/Order/WellFounded.lean +++ b/Mathlib/Order/WellFounded.lean @@ -171,7 +171,7 @@ theorem eq_strictMono_iff_eq_range {f g : β → γ} (hf : StrictMono f) (hg : S #align well_founded.eq_strict_mono_iff_eq_range WellFounded.eq_strictMono_iff_eq_range theorem self_le_of_strictMono {f : β → β} (hf : StrictMono f) : ∀ n, n ≤ f n := by - by_contra' h₁ + by_contra! h₁ have h₂ := h.min_mem _ h₁ exact h.not_lt_min _ h₁ (hf h₂) h₂ #align well_founded.self_le_of_strict_mono WellFounded.self_le_of_strictMono diff --git a/Mathlib/Order/WellFoundedSet.lean b/Mathlib/Order/WellFoundedSet.lean index 2ddea5912cf55..a9d586c15bee9 100644 --- a/Mathlib/Order/WellFoundedSet.lean +++ b/Mathlib/Order/WellFoundedSet.lean @@ -356,7 +356,7 @@ theorem partiallyWellOrderedOn_iff_finite_antichains [IsSymm α r] : s.PartiallyWellOrderedOn r ↔ ∀ t, t ⊆ s → IsAntichain r t → t.Finite := by refine' ⟨fun h t ht hrt => hrt.finite_of_partiallyWellOrderedOn (h.mono ht), _⟩ rintro hs f hf - by_contra' H + by_contra! H refine' infinite_range_of_injective (fun m n hmn => _) (hs _ (range_subset_iff.2 hf) _) · obtain h | h | h := lt_trichotomy m n · refine' (H _ _ h _).elim diff --git a/Mathlib/Probability/Cdf.lean b/Mathlib/Probability/Cdf.lean index d7a750fb55666..e2ed75947ff55 100644 --- a/Mathlib/Probability/Cdf.lean +++ b/Mathlib/Probability/Cdf.lean @@ -69,6 +69,7 @@ lemma tendsto_cdf_atBot : Tendsto (cdf μ) atBot (𝓝 0) := tendsto_condCdf_atB /-- The cdf tends to 1 at +∞. -/ lemma tendsto_cdf_atTop : Tendsto (cdf μ) atTop (𝓝 1) := tendsto_condCdf_atTop _ _ +set_option synthInstance.maxHeartbeats 25000 in lemma ofReal_cdf [IsProbabilityMeasure μ] (x : ℝ) : ENNReal.ofReal (cdf μ x) = μ (Iic x) := by have h := lintegral_condCdf ((Measure.dirac Unit.unit).prod μ) x simpa only [MeasureTheory.Measure.fst_prod, Measure.prod_prod, measure_univ, one_mul, diff --git a/Mathlib/Probability/Distributions/Exponential.lean b/Mathlib/Probability/Distributions/Exponential.lean index be513cc6f7e84..604f236b6ce4c 100644 --- a/Mathlib/Probability/Distributions/Exponential.lean +++ b/Mathlib/Probability/Distributions/Exponential.lean @@ -137,7 +137,7 @@ lemma lintegral_exponentialPdf_eq_one (r : ℝ) (hr : 0 < r) : ∫⁻ x, exponen exact set_lintegral_congr_fun isClosed_Ici.measurableSet (ae_of_all _ (fun x (hx : 0 ≤ x) ↦ exponentialPdf_of_nonneg hx)) simp only [leftSide, add_zero] - rw [rightSide, ENNReal.toReal_eq_one_iff, ←ENNReal.toReal_eq_one_iff] + rw [rightSide, ENNReal.toReal_eq_one_iff, ← ENNReal.toReal_eq_one_iff] rw [← integral_eq_lintegral_of_nonneg_ae (ae_of_all _ (fun _ ↦ by positivity))] · have IntegrOn : IntegrableOn (fun x ↦ rexp (-(r * x))) (Ioi 0) := by simp only [← neg_mul, exp_neg_integrableOn_Ioi 0 hr] @@ -199,7 +199,7 @@ lemma lintegral_exponentialPdf_eq_antiDeriv (r x : ℝ) (hr : 0 < r) : simp only [exponentialPdf_eq] rw[set_lintegral_congr_fun measurableSet_Icc (ae_of_all _ (by intro a ⟨(hle : _ ≤ a), _⟩; rw [if_pos hle]))] - rw [←ENNReal.toReal_eq_toReal _ ENNReal.ofReal_ne_top, ←integral_eq_lintegral_of_nonneg_ae + rw [← ENNReal.toReal_eq_toReal _ ENNReal.ofReal_ne_top, ← integral_eq_lintegral_of_nonneg_ae (eventually_of_forall fun _ ↦ le_of_lt (mul_pos hr (exp_pos _)))] have : ∫ a in uIoc 0 x, r * rexp (-(r * a)) = ∫ a in (0)..x, r * rexp (-(r * a)) := by rw [intervalIntegral.intervalIntegral_eq_integral_uIoc, smul_eq_mul, if_pos h, one_mul] diff --git a/Mathlib/Probability/Distributions/Gaussian.lean b/Mathlib/Probability/Distributions/Gaussian.lean index 672395b66b59d..d4c6425f6e0bf 100644 --- a/Mathlib/Probability/Distributions/Gaussian.lean +++ b/Mathlib/Probability/Distributions/Gaussian.lean @@ -95,7 +95,7 @@ lemma integrable_gaussianPdfReal (μ : ℝ) (v : ℝ≥0) : /-- The gaussian distribution pdf integrates to 1 when the variance is not zero. -/ lemma lintegral_gaussianPdfReal_eq_one (μ : ℝ) {v : ℝ≥0} (h : v ≠ 0) : ∫⁻ x, ENNReal.ofReal (gaussianPdfReal μ v x) = 1 := by - rw [←ENNReal.toReal_eq_one_iff] + rw [← ENNReal.toReal_eq_one_iff] have hfm : AEStronglyMeasurable (gaussianPdfReal μ v) volume := (stronglyMeasurable_gaussianPdfReal μ v).aestronglyMeasurable have hf : 0 ≤ₐₛ gaussianPdfReal μ v := ae_of_all _ (gaussianPdfReal_nonneg μ v) @@ -105,8 +105,8 @@ lemma lintegral_gaussianPdfReal_eq_one (μ : ℝ) {v : ℝ≥0} (h : v ≠ 0) : rw [integral_sub_right_eq_self (μ := volume) (fun a ↦ rexp (-a ^ 2 / ((2 : ℝ) * v))) μ] simp only [gt_iff_lt, zero_lt_two, zero_le_mul_right, ge_iff_le, div_eq_inv_mul, mul_inv_rev, mul_neg] - simp_rw [←neg_mul] - rw [neg_mul, integral_gaussian, ← Real.sqrt_inv, ←Real.sqrt_mul] + simp_rw [← neg_mul] + rw [neg_mul, integral_gaussian, ← Real.sqrt_inv, ← Real.sqrt_mul] · field_simp ring · positivity diff --git a/Mathlib/Probability/Independence/Basic.lean b/Mathlib/Probability/Independence/Basic.lean index b60aeb0f6de13..502124125e079 100644 --- a/Mathlib/Probability/Independence/Basic.lean +++ b/Mathlib/Probability/Independence/Basic.lean @@ -156,7 +156,7 @@ lemma iIndepSets.meas_biInter (h : iIndepSets π μ) (s : Finset ι) {f : ι → (iIndepSets_iff _ _).1 h s hf lemma iIndepSets.meas_iInter [Fintype ι] (h : iIndepSets π μ) (hs : ∀ i, s i ∈ π i) : - μ (⋂ i, s i) = ∏ i, μ (s i) := by simp [←h.meas_biInter _ fun _i _ ↦ hs _] + μ (⋂ i, s i) = ∏ i, μ (s i) := by simp [← h.meas_biInter _ fun _i _ ↦ hs _] set_option linter.uppercaseLean3 false in #align probability_theory.Indep_sets.meas_Inter ProbabilityTheory.iIndepSets.meas_iInter @@ -180,7 +180,7 @@ lemma iIndep.meas_biInter (hμ : iIndep m μ) (hs : ∀ i, i ∈ S → Measurabl μ (⋂ i ∈ S, s i) = ∏ i in S, μ (s i) := (iIndep_iff _ _).1 hμ _ hs lemma iIndep.meas_iInter [Fintype ι] (hμ : iIndep m μ) (hs : ∀ i, MeasurableSet[m i] (s i)) : - μ (⋂ i, s i) = ∏ i, μ (s i) := by simp [←hμ.meas_biInter fun _ _ ↦ hs _] + μ (⋂ i, s i) = ∏ i, μ (s i) := by simp [← hμ.meas_biInter fun _ _ ↦ hs _] lemma Indep_iff_IndepSets (m₁ m₂ : MeasurableSpace Ω) [MeasurableSpace Ω] (μ : Measure Ω) : Indep m₁ m₂ μ ↔ IndepSets {s | MeasurableSet[m₁] s} {s | MeasurableSet[m₂] s} μ := by diff --git a/Mathlib/Probability/Independence/Conditional.lean b/Mathlib/Probability/Independence/Conditional.lean index c26689e92c2c3..537198d6ff5d3 100644 --- a/Mathlib/Probability/Independence/Conditional.lean +++ b/Mathlib/Probability/Independence/Conditional.lean @@ -56,14 +56,14 @@ open scoped BigOperators MeasureTheory ENNReal namespace ProbabilityTheory -variable {Ω ι : Type _} +variable {Ω ι : Type*} section Definitions section variable (m' : MeasurableSpace Ω) - [mΩ : MeasurableSpace Ω] [TopologicalSpace Ω] [BorelSpace Ω] [PolishSpace Ω] [Nonempty Ω] + [mΩ : MeasurableSpace Ω] [StandardBorelSpace Ω] [Nonempty Ω] (hm' : m' ≤ mΩ) /-- A family of sets of sets `π : ι → Set (Set Ω)` is conditionally independent given `m'` with @@ -100,14 +100,13 @@ end `t₁ ∈ m₁, t₂ ∈ m₂`, `μ⟦t₁ ∩ t₂ | m'⟧ =ᵐ[μ] μ⟦t₁ | m'⟧ * μ⟦t₂ | m'⟧`. See `ProbabilityTheory.condIndep_iff`. -/ def CondIndep (m' m₁ m₂ : MeasurableSpace Ω) - [mΩ : MeasurableSpace Ω] [TopologicalSpace Ω] [BorelSpace Ω] [PolishSpace Ω] [Nonempty Ω] + [mΩ : MeasurableSpace Ω] [StandardBorelSpace Ω] [Nonempty Ω] (hm' : m' ≤ mΩ) (μ : Measure Ω := by volume_tac) [IsFiniteMeasure μ] : Prop := kernel.Indep m₁ m₂ (condexpKernel μ m') (μ.trim hm') section -variable (m' : MeasurableSpace Ω) - [mΩ : MeasurableSpace Ω] [TopologicalSpace Ω] [BorelSpace Ω] [PolishSpace Ω] [Nonempty Ω] +variable (m' : MeasurableSpace Ω) [mΩ : MeasurableSpace Ω] [StandardBorelSpace Ω] [Nonempty Ω] (hm' : m' ≤ mΩ) /-- A family of sets is conditionally independent if the family of measurable space structures they @@ -130,7 +129,7 @@ measurable space structures they generate on `Ω` is conditionally independent. with codomain having measurable space structure `m`, the generated measurable space structure is `m.comap g`. See `ProbabilityTheory.iCondIndepFun_iff`. -/ -def iCondIndepFun {β : ι → Type _} (m : ∀ x : ι, MeasurableSpace (β x)) +def iCondIndepFun {β : ι → Type*} (m : ∀ x : ι, MeasurableSpace (β x)) (f : ∀ x : ι, Ω → β x) (μ : Measure Ω := by volume_tac) [IsFiniteMeasure μ] : Prop := kernel.iIndepFun m f (condexpKernel μ m') (μ.trim hm') @@ -138,7 +137,7 @@ def iCondIndepFun {β : ι → Type _} (m : ∀ x : ι, MeasurableSpace (β x)) are conditionally independent. For a function `f` with codomain having measurable space structure `m`, the generated measurable space structure is `m.comap f`. See `ProbabilityTheory.condIndepFun_iff`. -/ -def CondIndepFun {β γ : Type _} [MeasurableSpace β] [MeasurableSpace γ] +def CondIndepFun {β γ : Type*} [MeasurableSpace β] [MeasurableSpace γ] (f : Ω → β) (g : Ω → γ) (μ : Measure Ω := by volume_tac) [IsFiniteMeasure μ] : Prop := kernel.IndepFun f g (condexpKernel μ m') (μ.trim hm') @@ -149,8 +148,7 @@ end Definitions section DefinitionLemmas section -variable (m' : MeasurableSpace Ω) - [mΩ : MeasurableSpace Ω] [TopologicalSpace Ω] [BorelSpace Ω] [PolishSpace Ω] [Nonempty Ω] +variable (m' : MeasurableSpace Ω) [mΩ : MeasurableSpace Ω] [StandardBorelSpace Ω] [Nonempty Ω] (hm' : m' ≤ mΩ) lemma iCondIndepSets_iff (π : ι → Set (Set Ω)) (hπ : ∀ i s (_hs : s ∈ π i), MeasurableSet s) @@ -269,15 +267,14 @@ end section CondIndep -lemma condIndep_iff_condIndepSets (m' m₁ m₂ : MeasurableSpace Ω) - [mΩ : MeasurableSpace Ω] [TopologicalSpace Ω] [BorelSpace Ω] [PolishSpace Ω] [Nonempty Ω] - (hm' : m' ≤ mΩ) (μ : Measure Ω ) [IsFiniteMeasure μ] : +lemma condIndep_iff_condIndepSets (m' m₁ m₂ : MeasurableSpace Ω) [mΩ : MeasurableSpace Ω] + [StandardBorelSpace Ω] [Nonempty Ω] (hm' : m' ≤ mΩ) (μ : Measure Ω ) [IsFiniteMeasure μ] : CondIndep m' m₁ m₂ hm' μ ↔ CondIndepSets m' hm' {s | MeasurableSet[m₁] s} {s | MeasurableSet[m₂] s} μ := by simp only [CondIndep, CondIndepSets, kernel.Indep] lemma condIndep_iff (m' m₁ m₂ : MeasurableSpace Ω) - [mΩ : MeasurableSpace Ω] [TopologicalSpace Ω] [BorelSpace Ω] [PolishSpace Ω] [Nonempty Ω] + [mΩ : MeasurableSpace Ω] [StandardBorelSpace Ω] [Nonempty Ω] (hm' : m' ≤ mΩ) (hm₁ : m₁ ≤ mΩ) (hm₂ : m₂ ≤ mΩ) (μ : Measure Ω) [IsFiniteMeasure μ] : CondIndep m' m₁ m₂ hm' μ ↔ ∀ t1 t2, MeasurableSet[m₁] t1 → MeasurableSet[m₂] t2 @@ -289,8 +286,7 @@ lemma condIndep_iff (m' m₁ m₂ : MeasurableSpace Ω) end CondIndep -variable (m' : MeasurableSpace Ω) - [mΩ : MeasurableSpace Ω] [TopologicalSpace Ω] [BorelSpace Ω] [PolishSpace Ω] [Nonempty Ω] +variable (m' : MeasurableSpace Ω) [mΩ : MeasurableSpace Ω] [StandardBorelSpace Ω] [Nonempty Ω] (hm' : m' ≤ mΩ) lemma iCondIndepSet_iff_iCondIndep (s : ι → Set Ω) (μ : Measure Ω) [IsFiniteMeasure μ] : @@ -322,14 +318,14 @@ lemma condIndepSet_iff (s t : Set Ω) (hs : MeasurableSet s) (ht : MeasurableSet CondIndepSet m' hm' s t μ ↔ (μ⟦s ∩ t | m'⟧) =ᵐ[μ] (μ⟦s | m'⟧) * (μ⟦t | m'⟧) := by rw [condIndepSet_iff_condIndepSets_singleton _ _ hs ht μ, condIndepSets_singleton_iff _ _ hs ht] -lemma iCondIndepFun_iff_iCondIndep {β : ι → Type _} +lemma iCondIndepFun_iff_iCondIndep {β : ι → Type*} (m : ∀ x : ι, MeasurableSpace (β x)) (f : ∀ x : ι, Ω → β x) (μ : Measure Ω) [IsFiniteMeasure μ] : iCondIndepFun m' hm' m f μ ↔ iCondIndep m' hm' (fun x ↦ MeasurableSpace.comap (f x) (m x)) μ := by simp only [iCondIndepFun, iCondIndep, kernel.iIndepFun] -lemma iCondIndepFun_iff {β : ι → Type _} +lemma iCondIndepFun_iff {β : ι → Type*} (m : ∀ x : ι, MeasurableSpace (β x)) (f : ∀ x : ι, Ω → β x) (hf : ∀ i, Measurable (f i)) (μ : Measure Ω) [IsFiniteMeasure μ] : iCondIndepFun m' hm' m f μ @@ -339,13 +335,13 @@ lemma iCondIndepFun_iff {β : ι → Type _} rw [iCondIndep_iff] exact fun i ↦ (hf i).comap_le -lemma condIndepFun_iff_condIndep {β γ : Type _} [mβ : MeasurableSpace β] +lemma condIndepFun_iff_condIndep {β γ : Type*} [mβ : MeasurableSpace β] [mγ : MeasurableSpace γ] (f : Ω → β) (g : Ω → γ) (μ : Measure Ω) [IsFiniteMeasure μ] : CondIndepFun m' hm' f g μ ↔ CondIndep m' (MeasurableSpace.comap f mβ) (MeasurableSpace.comap g mγ) hm' μ := by simp only [CondIndepFun, CondIndep, kernel.IndepFun] -lemma condIndepFun_iff {β γ : Type _} [mβ : MeasurableSpace β] [mγ : MeasurableSpace γ] +lemma condIndepFun_iff {β γ : Type*} [mβ : MeasurableSpace β] [mγ : MeasurableSpace γ] (f : Ω → β) (g : Ω → γ) (hf : Measurable f) (hg : Measurable g) (μ : Measure Ω) [IsFiniteMeasure μ] : CondIndepFun m' hm' f g μ ↔ ∀ t1 t2, MeasurableSet[MeasurableSpace.comap f mβ] t1 @@ -357,8 +353,7 @@ end DefinitionLemmas section CondIndepSets -variable {m' : MeasurableSpace Ω} - [mΩ : MeasurableSpace Ω] [TopologicalSpace Ω] [BorelSpace Ω] [PolishSpace Ω] [Nonempty Ω] +variable {m' : MeasurableSpace Ω} [mΩ : MeasurableSpace Ω] [StandardBorelSpace Ω] [Nonempty Ω] {hm' : m' ≤ mΩ} {μ : Measure Ω} [IsFiniteMeasure μ] @[symm] @@ -416,8 +411,7 @@ end CondIndepSets section CondIndepSet -variable {m' : MeasurableSpace Ω} - [mΩ : MeasurableSpace Ω] [TopologicalSpace Ω] [BorelSpace Ω] [PolishSpace Ω] [Nonempty Ω] +variable {m' : MeasurableSpace Ω} [mΩ : MeasurableSpace Ω] [StandardBorelSpace Ω] [Nonempty Ω] {hm' : m' ≤ mΩ} {μ : Measure Ω} [IsFiniteMeasure μ] theorem condIndepSet_empty_right (s : Set Ω) : CondIndepSet m' hm' s ∅ μ := @@ -431,34 +425,33 @@ end CondIndepSet section CondIndep @[symm] -theorem CondIndep.symm {m' m₁ m₂: MeasurableSpace Ω} - [mΩ : MeasurableSpace Ω] [TopologicalSpace Ω] [BorelSpace Ω] [PolishSpace Ω] [Nonempty Ω] - {hm' : m' ≤ mΩ} {μ : Measure Ω} [IsFiniteMeasure μ] +theorem CondIndep.symm {m' m₁ m₂: MeasurableSpace Ω} [mΩ : MeasurableSpace Ω] + [StandardBorelSpace Ω] [Nonempty Ω] {hm' : m' ≤ mΩ} {μ : Measure Ω} [IsFiniteMeasure μ] (h : CondIndep m' m₁ m₂ hm' μ) : CondIndep m' m₂ m₁ hm' μ := CondIndepSets.symm h theorem condIndep_bot_right (m₁ : MeasurableSpace Ω) {m' : MeasurableSpace Ω} - [mΩ : MeasurableSpace Ω] [TopologicalSpace Ω] [BorelSpace Ω] [PolishSpace Ω] [Nonempty Ω] + [mΩ : MeasurableSpace Ω] [StandardBorelSpace Ω] [Nonempty Ω] {hm' : m' ≤ mΩ} {μ : Measure Ω} [IsFiniteMeasure μ] : CondIndep m' m₁ ⊥ hm' μ := kernel.indep_bot_right m₁ theorem condIndep_bot_left (m₁ : MeasurableSpace Ω) {m' : MeasurableSpace Ω} - [mΩ : MeasurableSpace Ω] [TopologicalSpace Ω] [BorelSpace Ω] [PolishSpace Ω] [Nonempty Ω] + [mΩ : MeasurableSpace Ω] [StandardBorelSpace Ω] [Nonempty Ω] {hm' : m' ≤ mΩ} {μ : Measure Ω} [IsFiniteMeasure μ] : CondIndep m' ⊥ m₁ hm' μ := (kernel.indep_bot_right m₁).symm theorem condIndep_of_condIndep_of_le_left {m' m₁ m₂ m₃ : MeasurableSpace Ω} - [mΩ : MeasurableSpace Ω] [TopologicalSpace Ω] [BorelSpace Ω] [PolishSpace Ω] [Nonempty Ω] + [mΩ : MeasurableSpace Ω] [StandardBorelSpace Ω] [Nonempty Ω] {hm' : m' ≤ mΩ} {μ : Measure Ω} [IsFiniteMeasure μ] (h_indep : CondIndep m' m₁ m₂ hm' μ) (h31 : m₃ ≤ m₁) : CondIndep m' m₃ m₂ hm' μ := kernel.indep_of_indep_of_le_left h_indep h31 theorem condIndep_of_condIndep_of_le_right {m' m₁ m₂ m₃: MeasurableSpace Ω} - [mΩ : MeasurableSpace Ω] [TopologicalSpace Ω] [BorelSpace Ω] [PolishSpace Ω] [Nonempty Ω] + [mΩ : MeasurableSpace Ω] [StandardBorelSpace Ω] [Nonempty Ω] {hm' : m' ≤ mΩ} {μ : Measure Ω} [IsFiniteMeasure μ] (h_indep : CondIndep m' m₁ m₂ hm' μ) (h32 : m₃ ≤ m₂) : CondIndep m' m₁ m₃ hm' μ := @@ -472,7 +465,7 @@ end CondIndep section FromiCondIndepToCondIndep variable {m' : MeasurableSpace Ω} - [mΩ : MeasurableSpace Ω] [TopologicalSpace Ω] [BorelSpace Ω] [PolishSpace Ω] [Nonempty Ω] + [mΩ : MeasurableSpace Ω] [StandardBorelSpace Ω] [Nonempty Ω] {hm' : m' ≤ mΩ} {μ : Measure Ω} [IsFiniteMeasure μ] theorem iCondIndepSets.condIndepSets {s : ι → Set (Set Ω)} @@ -485,7 +478,7 @@ theorem iCondIndep.condIndep {m : ι → MeasurableSpace Ω} CondIndep m' (m i) (m j) hm' μ := kernel.iIndep.indep h_indep hij -theorem iCondIndepFun.condIndepFun {β : ι → Type _} +theorem iCondIndepFun.condIndepFun {β : ι → Type*} {m : ∀ x, MeasurableSpace (β x)} {f : ∀ i, Ω → β i} (hf_Indep : iCondIndepFun m' hm' m f μ) {i j : ι} (hij : i ≠ j) : CondIndepFun m' hm' (f i) (f j) μ := @@ -507,7 +500,7 @@ section FromMeasurableSpacesToSetsOfSets generating π-systems -/ variable {m' : MeasurableSpace Ω} - [mΩ : MeasurableSpace Ω] [TopologicalSpace Ω] [BorelSpace Ω] [PolishSpace Ω] [Nonempty Ω] + [mΩ : MeasurableSpace Ω] [StandardBorelSpace Ω] [Nonempty Ω] {hm' : m' ≤ mΩ} {μ : Measure Ω} [IsFiniteMeasure μ] theorem iCondIndep.iCondIndepSets {m : ι → MeasurableSpace Ω} @@ -528,8 +521,7 @@ section FromPiSystemsToMeasurableSpaces /-! ### Conditional independence of generating π-systems implies conditional independence of σ-algebras -/ -variable {m' m₁ m₂ : MeasurableSpace Ω} - [mΩ : MeasurableSpace Ω] [TopologicalSpace Ω] [BorelSpace Ω] [PolishSpace Ω] [Nonempty Ω] +variable {m' m₁ m₂ : MeasurableSpace Ω} [mΩ : MeasurableSpace Ω] [StandardBorelSpace Ω] [Nonempty Ω] {hm' : m' ≤ mΩ} {μ : Measure Ω} [IsFiniteMeasure μ] theorem CondIndepSets.condIndep @@ -617,8 +609,7 @@ section CondIndepSet -/ -variable {m' m₁ m₂ : MeasurableSpace Ω} - [mΩ : MeasurableSpace Ω] [TopologicalSpace Ω] [BorelSpace Ω] [PolishSpace Ω] [Nonempty Ω] +variable {m' m₁ m₂ : MeasurableSpace Ω} [mΩ : MeasurableSpace Ω] [StandardBorelSpace Ω] [Nonempty Ω] {hm' : m' ≤ mΩ} {s t : Set Ω} (S T : Set (Set Ω)) @@ -647,8 +638,8 @@ section CondIndepFun -/ -variable {β β' : Type _} {m' : MeasurableSpace Ω} - [mΩ : MeasurableSpace Ω] [TopologicalSpace Ω] [BorelSpace Ω] [PolishSpace Ω] [Nonempty Ω] +variable {β β' : Type*} {m' : MeasurableSpace Ω} + [mΩ : MeasurableSpace Ω] [StandardBorelSpace Ω] [Nonempty Ω] {hm' : m' ≤ mΩ} {μ : Measure Ω} [IsFiniteMeasure μ] {f : Ω → β} {g : Ω → β'} @@ -663,7 +654,7 @@ theorem condIndepFun_iff_condexp_inter_preimage_eq_mul {mβ : MeasurableSpace β · rintro ⟨s, hs, rfl⟩ ⟨t, ht, rfl⟩ exact h s t hs ht -theorem iCondIndepFun_iff_condexp_inter_preimage_eq_mul {β : ι → Type _} +theorem iCondIndepFun_iff_condexp_inter_preimage_eq_mul {β : ι → Type*} (m : ∀ x, MeasurableSpace (β x)) (f : ∀ i, Ω → β i) (hf : ∀ i, Measurable (f i)) : iCondIndepFun m' hm' m f μ ↔ ∀ (S : Finset ι) {sets : ∀ i : ι, Set (β i)} (_H : ∀ i, i ∈ S → MeasurableSet[m i] (sets i)), @@ -697,7 +688,7 @@ nonrec theorem CondIndepFun.symm {_ : MeasurableSpace β} {f g : Ω → β} CondIndepFun m' hm' g f μ := hfg.symm -theorem CondIndepFun.comp {γ γ' : Type _} {_mβ : MeasurableSpace β} {_mβ' : MeasurableSpace β'} +theorem CondIndepFun.comp {γ γ' : Type*} {_mβ : MeasurableSpace β} {_mβ' : MeasurableSpace β'} {_mγ : MeasurableSpace γ} {_mγ' : MeasurableSpace γ'} {φ : β → γ} {ψ : β' → γ'} (hfg : CondIndepFun m' hm' f g μ) (hφ : Measurable φ) (hψ : Measurable ψ) : CondIndepFun m' hm' (φ ∘ f) (ψ ∘ g) μ := @@ -706,13 +697,13 @@ theorem CondIndepFun.comp {γ γ' : Type _} {_mβ : MeasurableSpace β} {_mβ' : /-- If `f` is a family of mutually conditionally independent random variables (`iCondIndepFun m' hm' m f μ`) and `S, T` are two disjoint finite index sets, then the tuple formed by `f i` for `i ∈ S` is conditionally independent of the tuple `(f i)_i` for `i ∈ T`. -/ -theorem iCondIndepFun.condIndepFun_finset {β : ι → Type _} +theorem iCondIndepFun.condIndepFun_finset {β : ι → Type*} {m : ∀ i, MeasurableSpace (β i)} {f : ∀ i, Ω → β i} (S T : Finset ι) (hST : Disjoint S T) (hf_Indep : iCondIndepFun m' hm' m f μ) (hf_meas : ∀ i, Measurable (f i)) : CondIndepFun m' hm' (fun a (i : S) => f i a) (fun a (i : T) => f i a) μ := kernel.iIndepFun.indepFun_finset S T hST hf_Indep hf_meas -theorem iCondIndepFun.condIndepFun_prod {β : ι → Type _} +theorem iCondIndepFun.condIndepFun_prod {β : ι → Type*} {m : ∀ i, MeasurableSpace (β i)} {f : ∀ i, Ω → β i} (hf_Indep : iCondIndepFun m' hm' m f μ) (hf_meas : ∀ i, Measurable (f i)) (i j k : ι) (hik : i ≠ k) (hjk : j ≠ k) : CondIndepFun m' hm' (fun a => (f i a, f j a)) (f k) μ := diff --git a/Mathlib/Probability/Independence/ZeroOne.lean b/Mathlib/Probability/Independence/ZeroOne.lean index cadba2a9bfed8..ca8c36a61c8e1 100644 --- a/Mathlib/Probability/Independence/ZeroOne.lean +++ b/Mathlib/Probability/Independence/ZeroOne.lean @@ -28,7 +28,7 @@ open scoped MeasureTheory ENNReal namespace ProbabilityTheory -variable {α Ω ι : Type _} {_mα : MeasurableSpace α} {m m0 : MeasurableSpace Ω} +variable {α Ω ι : Type*} {_mα : MeasurableSpace α} {m m0 : MeasurableSpace Ω} {κ : kernel α Ω} {μα : Measure α} {μ : Measure Ω} theorem kernel.measure_eq_zero_or_one_or_top_of_indepSet_self {t : Set Ω} @@ -63,7 +63,7 @@ theorem measure_eq_zero_or_one_of_indepSet_self [IsFiniteMeasure μ] {t : Set Ω #align probability_theory.measure_eq_zero_or_one_of_indep_set_self ProbabilityTheory.measure_eq_zero_or_one_of_indepSet_self theorem condexp_eq_zero_or_one_of_condIndepSet_self - [TopologicalSpace Ω] [BorelSpace Ω] [PolishSpace Ω] [Nonempty Ω] + [StandardBorelSpace Ω] [Nonempty Ω] (hm : m ≤ m0) [hμ : IsFiniteMeasure μ] {t : Set Ω} (ht : MeasurableSet t) (h_indep : CondIndepSet m hm t t μ) : ∀ᵐ ω ∂μ, (μ⟦t | m⟧) ω = 0 ∨ (μ⟦t | m⟧) ω = 1 := by @@ -87,7 +87,7 @@ theorem indep_biSup_compl (h_le : ∀ n, s n ≤ m0) (h_indep : iIndep s μ) (t kernel.indep_biSup_compl h_le h_indep t #align probability_theory.indep_bsupr_compl ProbabilityTheory.indep_biSup_compl -theorem condIndep_biSup_compl [TopologicalSpace Ω] [BorelSpace Ω] [PolishSpace Ω] [Nonempty Ω] +theorem condIndep_biSup_compl [StandardBorelSpace Ω] [Nonempty Ω] (hm : m ≤ m0) [IsFiniteMeasure μ] (h_le : ∀ n, s n ≤ m0) (h_indep : iCondIndep m hm s μ) (t : Set ι) : CondIndep m (⨆ n ∈ t, s n) (⨆ n ∈ tᶜ, s n) hm μ := @@ -121,7 +121,7 @@ theorem indep_biSup_limsup (h_le : ∀ n, s n ≤ m0) (h_indep : iIndep s μ) (h kernel.indep_biSup_limsup h_le h_indep hf ht #align probability_theory.indep_bsupr_limsup ProbabilityTheory.indep_biSup_limsup -theorem condIndep_biSup_limsup [TopologicalSpace Ω] [BorelSpace Ω] [PolishSpace Ω] [Nonempty Ω] +theorem condIndep_biSup_limsup [StandardBorelSpace Ω] [Nonempty Ω] (hm : m ≤ m0) [IsFiniteMeasure μ] (h_le : ∀ n, s n ≤ m0) (h_indep : iCondIndep m hm s μ) (hf : ∀ t, p t → tᶜ ∈ f) {t : Set ι} (ht : p t) : @@ -147,7 +147,7 @@ theorem indep_iSup_directed_limsup (h_le : ∀ n, s n ≤ m0) (h_indep : iIndep kernel.indep_iSup_directed_limsup h_le h_indep hf hns hnsp #align probability_theory.indep_supr_directed_limsup ProbabilityTheory.indep_iSup_directed_limsup -theorem condIndep_iSup_directed_limsup [TopologicalSpace Ω] [BorelSpace Ω] [PolishSpace Ω] +theorem condIndep_iSup_directed_limsup [StandardBorelSpace Ω] [Nonempty Ω] (hm : m ≤ m0) [IsFiniteMeasure μ] (h_le : ∀ n, s n ≤ m0) (h_indep : iCondIndep m hm s μ) (hf : ∀ t, p t → tᶜ ∈ f) (hns : Directed (· ≤ ·) ns) (hnsp : ∀ a, p (ns a)) : @@ -173,7 +173,7 @@ theorem indep_iSup_limsup (h_le : ∀ n, s n ≤ m0) (h_indep : iIndep s μ) (hf kernel.indep_iSup_limsup h_le h_indep hf hns hnsp hns_univ #align probability_theory.indep_supr_limsup ProbabilityTheory.indep_iSup_limsup -theorem condIndep_iSup_limsup [TopologicalSpace Ω] [BorelSpace Ω] [PolishSpace Ω] [Nonempty Ω] +theorem condIndep_iSup_limsup [StandardBorelSpace Ω] [Nonempty Ω] (hm : m ≤ m0) [IsFiniteMeasure μ] (h_le : ∀ n, s n ≤ m0) (h_indep : iCondIndep m hm s μ) (hf : ∀ t, p t → tᶜ ∈ f) (hns : Directed (· ≤ ·) ns) (hnsp : ∀ a, p (ns a)) (hns_univ : ∀ n, ∃ a, n ∈ ns a) : @@ -192,7 +192,7 @@ theorem indep_limsup_self (h_le : ∀ n, s n ≤ m0) (h_indep : iIndep s μ) (hf kernel.indep_limsup_self h_le h_indep hf hns hnsp hns_univ #align probability_theory.indep_limsup_self ProbabilityTheory.indep_limsup_self -theorem condIndep_limsup_self [TopologicalSpace Ω] [BorelSpace Ω] [PolishSpace Ω] [Nonempty Ω] +theorem condIndep_limsup_self [StandardBorelSpace Ω] [Nonempty Ω] (hm : m ≤ m0) [IsFiniteMeasure μ] (h_le : ∀ n, s n ≤ m0) (h_indep : iCondIndep m hm s μ) (hf : ∀ t, p t → tᶜ ∈ f) (hns : Directed (· ≤ ·) ns) (hnsp : ∀ a, p (ns a)) (hns_univ : ∀ n, ∃ a, n ∈ ns a) : @@ -217,8 +217,8 @@ theorem measure_zero_or_one_of_measurableSet_limsup (h_le : ∀ n, s n ≤ m0) ( ht_tail #align probability_theory.measure_zero_or_one_of_measurable_set_limsup ProbabilityTheory.measure_zero_or_one_of_measurableSet_limsup -theorem condexp_zero_or_one_of_measurableSet_limsup [TopologicalSpace Ω] [BorelSpace Ω] - [PolishSpace Ω] [Nonempty Ω] (hm : m ≤ m0) [IsFiniteMeasure μ] +theorem condexp_zero_or_one_of_measurableSet_limsup [StandardBorelSpace Ω] [Nonempty Ω] + (hm : m ≤ m0) [IsFiniteMeasure μ] (h_le : ∀ n, s n ≤ m0) (h_indep : iCondIndep m hm s μ) (hf : ∀ t, p t → tᶜ ∈ f) (hns : Directed (· ≤ ·) ns) (hnsp : ∀ a, p (ns a)) (hns_univ : ∀ n, ∃ a, n ∈ ns a) {t : Set Ω} (ht_tail : MeasurableSet[limsup s f] t) : @@ -257,7 +257,7 @@ theorem indep_limsup_atTop_self (h_le : ∀ n, s n ≤ m0) (h_indep : iIndep s kernel.indep_limsup_atTop_self h_le h_indep #align probability_theory.indep_limsup_at_top_self ProbabilityTheory.indep_limsup_atTop_self -theorem condIndep_limsup_atTop_self [TopologicalSpace Ω] [BorelSpace Ω] [PolishSpace Ω] [Nonempty Ω] +theorem condIndep_limsup_atTop_self [StandardBorelSpace Ω] [Nonempty Ω] (hm : m ≤ m0) [IsFiniteMeasure μ] (h_le : ∀ n, s n ≤ m0) (h_indep : iCondIndep m hm s μ) : CondIndep m (limsup s atTop) (limsup s atTop) hm μ := @@ -279,9 +279,8 @@ theorem measure_zero_or_one_of_measurableSet_limsup_atTop (h_le : ∀ n, s n ≤ using kernel.measure_zero_or_one_of_measurableSet_limsup_atTop h_le h_indep ht_tail #align probability_theory.measure_zero_or_one_of_measurable_set_limsup_at_top ProbabilityTheory.measure_zero_or_one_of_measurableSet_limsup_atTop -theorem condexp_zero_or_one_of_measurableSet_limsup_atTop [TopologicalSpace Ω] [BorelSpace Ω] - [PolishSpace Ω] [Nonempty Ω] (hm : m ≤ m0) [IsFiniteMeasure μ] - (h_le : ∀ n, s n ≤ m0) +theorem condexp_zero_or_one_of_measurableSet_limsup_atTop [StandardBorelSpace Ω] [Nonempty Ω] + (hm : m ≤ m0) [IsFiniteMeasure μ] (h_le : ∀ n, s n ≤ m0) (h_indep : iCondIndep m hm s μ) {t : Set Ω} (ht_tail : MeasurableSet[limsup s atTop] t) : ∀ᵐ ω ∂μ, (μ⟦t | m⟧) ω = 0 ∨ (μ⟦t | m⟧) ω = 1 := condexp_eq_zero_or_one_of_condIndepSet_self hm (limsup_le_iSup.trans (iSup_le h_le) t ht_tail) @@ -312,7 +311,7 @@ theorem indep_limsup_atBot_self (h_le : ∀ n, s n ≤ m0) (h_indep : iIndep s kernel.indep_limsup_atBot_self h_le h_indep #align probability_theory.indep_limsup_at_bot_self ProbabilityTheory.indep_limsup_atBot_self -theorem condIndep_limsup_atBot_self [TopologicalSpace Ω] [BorelSpace Ω] [PolishSpace Ω] [Nonempty Ω] +theorem condIndep_limsup_atBot_self [StandardBorelSpace Ω] [Nonempty Ω] (hm : m ≤ m0) [IsFiniteMeasure μ] (h_le : ∀ n, s n ≤ m0) (h_indep : iCondIndep m hm s μ) : CondIndep m (limsup s atBot) (limsup s atBot) hm μ := @@ -337,9 +336,8 @@ theorem measure_zero_or_one_of_measurableSet_limsup_atBot (h_le : ∀ n, s n ≤ /-- **Kolmogorov's 0-1 law**, conditional version: any event in the tail σ-algebra of a conditinoally independent sequence of sub-σ-algebras has conditional probability 0 or 1. -/ -theorem condexp_zero_or_one_of_measurableSet_limsup_atBot [TopologicalSpace Ω] [BorelSpace Ω] - [PolishSpace Ω] [Nonempty Ω] (hm : m ≤ m0) [IsFiniteMeasure μ] - (h_le : ∀ n, s n ≤ m0) +theorem condexp_zero_or_one_of_measurableSet_limsup_atBot [StandardBorelSpace Ω] [Nonempty Ω] + (hm : m ≤ m0) [IsFiniteMeasure μ] (h_le : ∀ n, s n ≤ m0) (h_indep : iCondIndep m hm s μ) {t : Set Ω} (ht_tail : MeasurableSet[limsup s atBot] t) : ∀ᵐ ω ∂μ, (μ⟦t | m⟧) ω = 0 ∨ (μ⟦t | m⟧) ω = 1 := condexp_eq_zero_or_one_of_condIndepSet_self hm (limsup_le_iSup.trans (iSup_le h_le) t ht_tail) diff --git a/Mathlib/Probability/Kernel/Basic.lean b/Mathlib/Probability/Kernel/Basic.lean index 0933d7c6bb701..2aee81274052f 100644 --- a/Mathlib/Probability/Kernel/Basic.lean +++ b/Mathlib/Probability/Kernel/Basic.lean @@ -97,7 +97,7 @@ theorem zero_apply (a : α) : (0 : kernel α β) a = 0 := @[simp] theorem coe_finset_sum (I : Finset ι) (κ : ι → kernel α β) : ⇑(∑ i in I, κ i) = ∑ i in I, ⇑(κ i) := - (coeAddHom α β).map_sum _ _ + map_sum (coeAddHom α β) _ _ #align probability_theory.kernel.coe_finset_sum ProbabilityTheory.kernel.coe_finset_sum theorem finset_sum_apply (I : Finset ι) (κ : ι → kernel α β) (a : α) : diff --git a/Mathlib/RepresentationTheory/Basic.lean b/Mathlib/RepresentationTheory/Basic.lean index 27d6ea2fdc076..ffb38cb0bc44e 100644 --- a/Mathlib/RepresentationTheory/Basic.lean +++ b/Mathlib/RepresentationTheory/Basic.lean @@ -361,7 +361,7 @@ theorem ofMulAction_self_smul_eq_mul (x : MonoidAlgebra k G) (y : (ofMulAction k (fun x y hx hy => by simp only [hx, hy, add_mul, add_smul]) fun r x hx => by show asAlgebraHom (ofMulAction k G G) _ _ = _ -- Porting note: was simpa [← hx] simp only [map_smul, smul_apply, Algebra.smul_mul_assoc] - rw [←hx] + rw [← hx] rfl #align representation.of_mul_action_self_smul_eq_mul Representation.ofMulAction_self_smul_eq_mul diff --git a/Mathlib/RepresentationTheory/GroupCohomology/Basic.lean b/Mathlib/RepresentationTheory/GroupCohomology/Basic.lean index 9bbc484be6b30..c2e8043a3b16c 100644 --- a/Mathlib/RepresentationTheory/GroupCohomology/Basic.lean +++ b/Mathlib/RepresentationTheory/GroupCohomology/Basic.lean @@ -135,7 +135,7 @@ and the homogeneous `linearYonedaObjResolution`. -/ LinearEquiv.toModuleIso_inv, linearYonedaObjResolution_d_apply, LinearEquiv.toModuleIso_hom, diagonalHomEquiv_apply, Action.comp_hom, Resolution.d_eq k G n, Resolution.d_of (Fin.partialProd g), LinearMap.map_sum, - ←Finsupp.smul_single_one _ ((-1 : k) ^ _), map_smul, d_apply] + ← Finsupp.smul_single_one _ ((-1 : k) ^ _), map_smul, d_apply] simp only [@Fin.sum_univ_succ _ _ (n + 1), Fin.val_zero, pow_zero, one_smul, Fin.succAbove_zero, diagonalHomEquiv_symm_apply f (Fin.partialProd g ∘ @Fin.succ (n + 1)), Function.comp_apply, Fin.partialProd_succ, Fin.castSucc_zero, Fin.partialProd_zero, one_mul] @@ -155,7 +155,7 @@ and the homogeneous `linearYonedaObjResolution`. -/ erw [diagonalHomEquiv_apply, Action.comp_hom, ModuleCat.comp_def, LinearMap.comp_apply, resolution.d_eq] erw [resolution.d_of (Fin.partialProd g)] - simp only [map_sum, ←Finsupp.smul_single_one _ ((-1 : k) ^ _)] + simp only [map_sum, ← Finsupp.smul_single_one _ ((-1 : k) ^ _)] -- This used to be `rw`, but we need `erw` after leanprover/lean4#2644 erw [d_apply, @Fin.sum_univ_succ _ _ (n + 1), Fin.val_zero, pow_zero, one_smul, Fin.succAbove_zero, diagonalHomEquiv_symm_apply f (Fin.partialProd g ∘ @Fin.succ (n + 1))] diff --git a/Mathlib/RepresentationTheory/GroupCohomology/Hilbert90.lean b/Mathlib/RepresentationTheory/GroupCohomology/Hilbert90.lean index 308387d24eb5a..31882fc1b2317 100644 --- a/Mathlib/RepresentationTheory/GroupCohomology/Hilbert90.lean +++ b/Mathlib/RepresentationTheory/GroupCohomology/Hilbert90.lean @@ -83,7 +83,7 @@ theorem hilbert90 (f : (L ≃ₐ[K] L) → Lˣ) intro g /- Then the equality follows from the hypothesis `hf` (that `f` is a 1-cocycle). -/ simp_rw [Units.ext_iff, this, Units.val_mul, Units.coe_map, Units.val_mk0, MonoidHom.coe_coe, - map_sum, map_mul, Finset.mul_sum, ←mul_assoc, mul_comm (f _ : L), ←hf, mul_comm (f _ : L)] + map_sum, map_mul, Finset.mul_sum, ← mul_assoc, mul_comm (f _ : L), ← hf, mul_comm (f _ : L)] exact Fintype.sum_bijective (fun i => g * i) (Group.mulLeft_bijective g) _ _ (fun i => rfl) end diff --git a/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean b/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean index b2e6d1ff3f185..bd779ec3f2dce 100644 --- a/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean +++ b/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean @@ -196,7 +196,7 @@ theorem dTwo_comp_dOne : dTwo A ∘ₗ dOne A = 0 := by show ModuleCat.ofHom (dOne A) ≫ ModuleCat.ofHom (dTwo A) = _ have h1 : _ ≫ ModuleCat.ofHom (dOne A) = _ ≫ _ := congr_arg ModuleCat.ofHom (dOne_comp_eq A) have h2 : _ ≫ ModuleCat.ofHom (dTwo A) = _ ≫ _ := congr_arg ModuleCat.ofHom (dTwo_comp_eq A) - simp only [←LinearEquiv.toModuleIso_hom] at h1 h2 + simp only [← LinearEquiv.toModuleIso_hom] at h1 h2 simp only [(Iso.eq_inv_comp _).2 h2, (Iso.eq_inv_comp _).2 h1, Category.assoc, Iso.hom_inv_id_assoc, HomologicalComplex.d_comp_d_assoc, zero_comp, comp_zero] diff --git a/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean b/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean index afec0afc4a029..292077fcfd62d 100644 --- a/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean +++ b/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean @@ -149,7 +149,7 @@ theorem actionDiagonalSucc_inv_apply {G : Type u} [Group G] {n : ℕ} (g : G) (f erw [hn, Equiv.piFinSuccAboveEquiv_symm_apply] refine' Fin.cases _ (fun i => _) x · simp only [Fin.insertNth_zero, Fin.cons_zero, Fin.partialProd_zero, mul_one] - · simp only [Fin.cons_succ, Pi.smul_apply, smul_eq_mul, Fin.partialProd_succ', ←mul_assoc] + · simp only [Fin.cons_succ, Pi.smul_apply, smul_eq_mul, Fin.partialProd_succ', ← mul_assoc] rfl set_option linter.uppercaseLean3 false in #align group_cohomology.resolution.Action_diagonal_succ_inv_apply groupCohomology.resolution.actionDiagonalSucc_inv_apply @@ -290,7 +290,7 @@ def ofMulActionBasisAux : (Rep.trivial k G ((Fin n → G) →₀ k))).ρ r _ refine' x.induction_on _ (fun x y => _) fun y z hy hz => _ · rw [smul_zero, map_zero] - · rw [TensorProduct.smul_tmul', smul_eq_mul, ←ofMulAction_self_smul_eq_mul] + · rw [TensorProduct.smul_tmul', smul_eq_mul, ← ofMulAction_self_smul_eq_mul] exact (smul_tprod_one_asModule (Representation.ofMulAction k G G) r x y).symm · rw [smul_add, hz, hy, map_add] } #align group_cohomology.resolution.of_mul_action_basis_aux groupCohomology.resolution.ofMulActionBasisAux @@ -655,7 +655,7 @@ theorem d_comp_ε : (groupCohomology.resolution k G).d 1 0 ≫ ε k G = 0 := by have : (forget₂ToModuleCat k G).d 1 0 ≫ (forget₂ (Rep k G) (ModuleCat.{u} k)).map (ε k G) = 0 := by rw [← forget₂ToModuleCatHomotopyEquiv_f_0_eq, - ←(forget₂ToModuleCatHomotopyEquiv k G).1.2 1 0 rfl] + ← (forget₂ToModuleCatHomotopyEquiv k G).1.2 1 0 rfl] exact comp_zero exact LinearMap.ext_iff.1 this _ #align group_cohomology.resolution.d_comp_ε groupCohomology.resolution.d_comp_ε diff --git a/Mathlib/RepresentationTheory/Rep.lean b/Mathlib/RepresentationTheory/Rep.lean index 8ca492a07c15f..0fa3e7d79be5b 100644 --- a/Mathlib/RepresentationTheory/Rep.lean +++ b/Mathlib/RepresentationTheory/Rep.lean @@ -372,7 +372,7 @@ noncomputable def leftRegularHomEquiv (A : Rep k G) : (Rep.ofMulAction k G G ⟶ Representation.ofMulAction_single x (1 : G) (1 : k), smul_eq_mul, mul_one] -/ simp only [LinearMap.comp_apply, Finsupp.lsingle_apply, leftRegularHom_hom] erw [Finsupp.lift_apply] - rw [Finsupp.sum_single_index, ←this, of_ρ_apply] + rw [Finsupp.sum_single_index, ← this, of_ρ_apply] erw [Representation.ofMulAction_single x (1 : G) (1 : k)] simp only [one_smul, smul_eq_mul, mul_one] · -- This goal didn't exist before leanprover/lean4#2644 @@ -431,9 +431,9 @@ def homEquiv (A B C : Rep k G) : (A ⊗ B ⟶ C) ≃ (B ⟶ (Rep.ihom A).obj C) comm := fun g => by refine' LinearMap.ext fun x => LinearMap.ext fun y => _ change f.hom (_ ⊗ₜ[k] _) = C.ρ g (f.hom (_ ⊗ₜ[k] _)) - rw [←hom_comm_apply] + rw [← hom_comm_apply] change _ = f.hom ((A.ρ g * A.ρ g⁻¹) y ⊗ₜ[k] _) - simp only [←map_mul, mul_inv_self, map_one] + simp only [← map_mul, mul_inv_self, map_one] rfl } invFun f := { hom := TensorProduct.uncurry k _ _ _ f.hom.flip diff --git a/Mathlib/RingTheory/DedekindDomain/Ideal.lean b/Mathlib/RingTheory/DedekindDomain/Ideal.lean index 27aae3ac721da..02582c82b872d 100644 --- a/Mathlib/RingTheory/DedekindDomain/Ideal.lean +++ b/Mathlib/RingTheory/DedekindDomain/Ideal.lean @@ -298,7 +298,7 @@ theorem integrallyClosed : IsIntegrallyClosed A := by -- `A[x]` (which is a fractional ideal) is in fact equal to `A`. refine ⟨fun {x hx} => ?_⟩ rw [← Set.mem_range, ← Algebra.mem_bot, ← Subalgebra.mem_toSubmodule, Algebra.toSubmodule_bot, - Submodule.one_eq_span, ←coe_spanSingleton A⁰ (1 : FractionRing A), spanSingleton_one, ← + Submodule.one_eq_span, ← coe_spanSingleton A⁰ (1 : FractionRing A), spanSingleton_one, ← FractionalIdeal.adjoinIntegral_eq_one_of_isUnit x hx (h.isUnit _)] · exact mem_adjoinIntegral_self A⁰ x hx · exact fun h => one_ne_zero (eq_zero_iff.mp h 1 (Algebra.adjoin A {x}).one_mem) diff --git a/Mathlib/RingTheory/DedekindDomain/PID.lean b/Mathlib/RingTheory/DedekindDomain/PID.lean index ca225eb4dc8a9..b38dbdc23d093 100644 --- a/Mathlib/RingTheory/DedekindDomain/PID.lean +++ b/Mathlib/RingTheory/DedekindDomain/PID.lean @@ -125,7 +125,7 @@ theorem FractionalIdeal.isPrincipal.of_finite_maximals_of_inv {A : Type*} [CommR left_lt_sup.1 ((hf.mem_toFinset.1 hM).ne_top.lt_top.trans_eq (Ideal.sup_iInf_eq_top <| coprime M hM).symm) have : ∀ M ∈ s, ∃ a ∈ I, ∃ b ∈ I', a * b ∉ IsLocalization.coeSubmodule A M := by - intro M hM; by_contra' h + intro M hM; by_contra! h obtain ⟨x, hx, hxM⟩ := SetLike.exists_of_lt ((IsLocalization.coeSubmodule_strictMono hS (hf.mem_toFinset.1 hM).ne_top.lt_top).trans_eq diff --git a/Mathlib/RingTheory/DiscreteValuationRing/Basic.lean b/Mathlib/RingTheory/DiscreteValuationRing/Basic.lean index 1c9365020c8f3..e7f9830fd3e76 100644 --- a/Mathlib/RingTheory/DiscreteValuationRing/Basic.lean +++ b/Mathlib/RingTheory/DiscreteValuationRing/Basic.lean @@ -5,8 +5,7 @@ Authors: Kevin Buzzard -/ import Mathlib.RingTheory.PrincipalIdealDomain import Mathlib.RingTheory.Ideal.LocalRing -import Mathlib.RingTheory.Multiplicity -import Mathlib.RingTheory.Valuation.Basic +import Mathlib.RingTheory.Valuation.PrimeMultiplicity import Mathlib.LinearAlgebra.AdicCompletion #align_import ring_theory.discrete_valuation_ring.basic from "leanprover-community/mathlib"@"c163ec99dfc664628ca15d215fce0a5b9c265b68" @@ -78,7 +77,7 @@ theorem irreducible_of_span_eq_maximalIdeal {R : Type*} [CommRing R] [LocalRing have h2 : ¬IsUnit ϖ := show ϖ ∈ maximalIdeal R from h.symm ▸ Submodule.mem_span_singleton_self ϖ refine' ⟨h2, _⟩ intro a b hab - by_contra' h + by_contra! h obtain ⟨ha : a ∈ maximalIdeal R, hb : b ∈ maximalIdeal R⟩ := h rw [h, mem_span_singleton'] at ha hb rcases ha with ⟨a, rfl⟩ diff --git a/Mathlib/RingTheory/DiscreteValuationRing/TFAE.lean b/Mathlib/RingTheory/DiscreteValuationRing/TFAE.lean index 4a9537149e0a6..c9ae1e0efb3de 100644 --- a/Mathlib/RingTheory/DiscreteValuationRing/TFAE.lean +++ b/Mathlib/RingTheory/DiscreteValuationRing/TFAE.lean @@ -68,15 +68,14 @@ theorem exists_maximalIdeal_pow_eq_of_principal [IsNoetherianRing R] [LocalRing exact Associated.mul_mul (this _ (Multiset.mem_cons_self _ _)) hn have : ∃ n : ℕ, x ^ n ∈ I := by obtain ⟨r, hr₁, hr₂⟩ : ∃ r : R, r ∈ I ∧ r ≠ 0 := by - by_contra' h; apply hI; rw [eq_bot_iff]; exact h + by_contra! h; apply hI; rw [eq_bot_iff]; exact h obtain ⟨n, u, rfl⟩ := H' r hr₂ (le_maximalIdeal hI' hr₁) use n rwa [← I.unit_mul_mem_iff_mem u.isUnit, mul_comm] use Nat.find this apply le_antisymm · change ∀ s ∈ I, s ∈ _ - by_contra hI'' - push_neg at hI'' + by_contra! hI'' obtain ⟨s, hs₁, hs₂⟩ := hI'' apply hs₂ by_cases hs₃ : s = 0; · rw [hs₃]; exact zero_mem _ @@ -95,7 +94,7 @@ theorem maximalIdeal_isPrincipal_of_isDedekindDomain [LocalRing R] [IsDomain R] by_cases ne_bot : maximalIdeal R = ⊥ · rw [ne_bot]; infer_instance obtain ⟨a, ha₁, ha₂⟩ : ∃ a ∈ maximalIdeal R, a ≠ (0 : R) := by - by_contra' h'; apply ne_bot; rwa [eq_bot_iff] + by_contra! h'; apply ne_bot; rwa [eq_bot_iff] have hle : Ideal.span {a} ≤ maximalIdeal R := by rwa [Ideal.span_le, Set.singleton_subset_iff] have : (Ideal.span {a}).radical = maximalIdeal R := by rw [Ideal.radical_eq_sInf] @@ -112,7 +111,7 @@ theorem maximalIdeal_isPrincipal_of_isDedekindDomain [LocalRing R] [IsDomain R] rw [hn, pow_zero, Ideal.one_eq_top] at this exact (Ideal.IsMaximal.ne_top inferInstance (eq_top_iff.mpr <| this.trans hle)).elim obtain ⟨b, hb₁, hb₂⟩ : ∃ b ∈ maximalIdeal R ^ n, ¬b ∈ Ideal.span {a} := by - by_contra' h'; rw [Nat.find_eq_iff] at hn; exact hn.2 n n.lt_succ_self fun x hx => h' x hx + by_contra! h'; rw [Nat.find_eq_iff] at hn; exact hn.2 n n.lt_succ_self fun x hx => h' x hx have hb₃ : ∀ m ∈ maximalIdeal R, ∃ k : R, k * a = b * m := by intro m hm; rw [← Ideal.mem_span_singleton']; apply Nat.find_spec this rw [hn, pow_succ']; exact Ideal.mul_mem_mul hb₁ hm diff --git a/Mathlib/RingTheory/FiniteType.lean b/Mathlib/RingTheory/FiniteType.lean index 3389f8d56c282..c14c1cba38d24 100644 --- a/Mathlib/RingTheory/FiniteType.lean +++ b/Mathlib/RingTheory/FiniteType.lean @@ -722,7 +722,7 @@ theorem Module.Finite.injective_of_surjective_endomorphism {R : Type*} [CommRing have : (⊤ : Submodule R[X] (AEval' f)) ≤ Ideal.span {(X : R[X])} • ⊤ · intro a _ obtain ⟨y, rfl⟩ := f_surj.comp (AEval'.of f).symm.surjective a - rw [Function.comp_apply, ←AEval'.of_symm_X_smul] + rw [Function.comp_apply, ← AEval'.of_symm_X_smul] exact Submodule.smul_mem_smul (Ideal.mem_span_singleton.mpr (dvd_refl _)) trivial obtain ⟨F, hFa, hFb⟩ := Submodule.exists_sub_one_mem_and_smul_eq_zero_of_fg_of_le_smul _ (⊤ : Submodule R[X] (AEval' f)) diff --git a/Mathlib/RingTheory/Flat.lean b/Mathlib/RingTheory/Flat.lean index 694226c68a72c..80ce3c049cbd1 100644 --- a/Mathlib/RingTheory/Flat.lean +++ b/Mathlib/RingTheory/Flat.lean @@ -141,7 +141,7 @@ lemma of_retract [f : Flat R M] (i : N →ₗ[R] M) (r : M →ₗ[R] N) (h : r.c rw [← LinearMap.comp_apply, ← lTensor_comp, h] simp rw [← Function.Injective.of_comp_iff h₁ (rTensor N I.subtype), ← LinearMap.coe_comp] - rw [LinearMap.lTensor_comp_rTensor, ←LinearMap.rTensor_comp_lTensor] + rw [LinearMap.lTensor_comp_rTensor, ← LinearMap.rTensor_comp_lTensor] rw [LinearMap.coe_comp, Function.Injective.of_comp_iff (f hI)] apply Function.RightInverse.injective (g := lTensor _ r) intro x diff --git a/Mathlib/RingTheory/GradedAlgebra/Radical.lean b/Mathlib/RingTheory/GradedAlgebra/Radical.lean index fd7b5d5d6601c..280da6b16972a 100644 --- a/Mathlib/RingTheory/GradedAlgebra/Radical.lean +++ b/Mathlib/RingTheory/GradedAlgebra/Radical.lean @@ -56,8 +56,7 @@ theorem Ideal.IsHomogeneous.isPrime_of_homogeneous_mem_or_mem {I : Ideal A} (hI Ideal.IsPrime I := ⟨I_ne_top, by intro x y hxy - by_contra rid - push_neg at rid + by_contra! rid obtain ⟨rid₁, rid₂⟩ := rid classical /- diff --git a/Mathlib/RingTheory/HahnSeries.lean b/Mathlib/RingTheory/HahnSeries.lean index a2d004d0fc3b6..e61875f401cdd 100644 --- a/Mathlib/RingTheory/HahnSeries.lean +++ b/Mathlib/RingTheory/HahnSeries.lean @@ -798,7 +798,7 @@ private theorem mul_assoc' [NonUnitalSemiring R] (x y z : HahnSeries Γ R) : obtain ⟨⟨nx, H, rfl⟩, ny, nz, rfl⟩ := H1 exact ⟨i + k, l, i, k, ⟨⟨Set.add_mem_add nx ny, nz, add_assoc _ _ _⟩ , nx, ny, rfl⟩, - fun h => H2 <| by rw [←h, mul_assoc], rfl⟩ + fun h => H2 <| by rw [← h, mul_assoc], rfl⟩ · rintro ⟨⟨i, j⟩, ⟨k, l⟩⟩ _ _ simp [mul_assoc] diff --git a/Mathlib/RingTheory/Ideal/Basic.lean b/Mathlib/RingTheory/Ideal/Basic.lean index c1a345809478f..410e72ebb6203 100644 --- a/Mathlib/RingTheory/Ideal/Basic.lean +++ b/Mathlib/RingTheory/Ideal/Basic.lean @@ -650,7 +650,7 @@ lemma isPrime_of_maximally_disjoint (I : Ideal α) have : 1 ∈ (S : Set α) := S.one_mem aesop mem_or_mem' {x y} hxy := by - by_contra' rid + by_contra! rid have hx := maximally_disjoint (I ⊔ span {x}) (Submodule.lt_sup_iff_not_mem.mpr rid.1) have hy := maximally_disjoint (I ⊔ span {y}) (Submodule.lt_sup_iff_not_mem.mpr rid.2) simp only [Set.not_disjoint_iff, mem_inter_iff, SetLike.mem_coe, Submodule.mem_sup, @@ -844,7 +844,7 @@ theorem bot_lt_of_maximal (M : Ideal R) [hm : M.IsMaximal] (non_field : ¬IsFiel intro mle apply lt_irrefl (⊤ : Ideal R) have : M = ⊥ := eq_bot_iff.mpr mle - rw [←this] at Ibot + rw [← this] at Ibot rwa [hm.1.2 I Ibot] at Itop #align ideal.bot_lt_of_maximal Ideal.bot_lt_of_maximal diff --git a/Mathlib/RingTheory/Ideal/Over.lean b/Mathlib/RingTheory/Ideal/Over.lean index 84fb48209fd96..c059a2e2d5cc1 100644 --- a/Mathlib/RingTheory/Ideal/Over.lean +++ b/Mathlib/RingTheory/Ideal/Over.lean @@ -370,7 +370,7 @@ theorem exists_ideal_over_prime_of_isIntegral_of_isDomain (H : Algebra.IsIntegra refine' ⟨comap (algebraMap S Sₚ) Qₚ, ⟨comap_isPrime _ Qₚ, _⟩⟩ convert Localization.AtPrime.comap_maximalIdeal (I := P) rw [comap_comap, ← LocalRing.eq_maximalIdeal Qₚ_max, - ←@IsLocalization.map_comp (P := S) (Q := Sₚ) (g := algebraMap R S) + ← @IsLocalization.map_comp (P := S) (Q := Sₚ) (g := algebraMap R S) (M := P.primeCompl) (T := Algebra.algebraMapSubmonoid S P.primeCompl) (S := Rₚ) _ _ _ _ _ _ (fun p hp => Algebra.mem_algebraMapSubmonoid_of_mem ⟨p, hp⟩) _ _] rfl diff --git a/Mathlib/RingTheory/Int/Basic.lean b/Mathlib/RingTheory/Int/Basic.lean index ba62b5f36f031..c616f9e1b1eb4 100644 --- a/Mathlib/RingTheory/Int/Basic.lean +++ b/Mathlib/RingTheory/Int/Basic.lean @@ -265,7 +265,7 @@ def associatesIntEquivNat : Associates ℤ ≃ ℕ := by theorem Int.Prime.dvd_mul {m n : ℤ} {p : ℕ} (hp : Nat.Prime p) (h : (p : ℤ) ∣ m * n) : p ∣ m.natAbs ∨ p ∣ n.natAbs := by - rwa [← hp.dvd_mul, ← Int.natAbs_mul, ←Int.coe_nat_dvd_left] + rwa [← hp.dvd_mul, ← Int.natAbs_mul, ← Int.coe_nat_dvd_left] #align int.prime.dvd_mul Int.Prime.dvd_mul theorem Int.Prime.dvd_mul' {m n : ℤ} {p : ℕ} (hp : Nat.Prime p) (h : (p : ℤ) ∣ m * n) : diff --git a/Mathlib/RingTheory/IntegralClosure.lean b/Mathlib/RingTheory/IntegralClosure.lean index 087a14337447b..2cc38fdf8f000 100644 --- a/Mathlib/RingTheory/IntegralClosure.lean +++ b/Mathlib/RingTheory/IntegralClosure.lean @@ -267,8 +267,7 @@ theorem isIntegral_of_smul_mem_submodule {M : Type*} [AddCommGroup M] [Module R (by ext; apply one_smul) (by intros x y; ext; apply mul_smul) obtain ⟨a, ha₁, ha₂⟩ : ∃ a ∈ N, a ≠ (0 : M) := by - by_contra h' - push_neg at h' + by_contra! h' apply hN rwa [eq_bot_iff] have : Function.Injective f := by diff --git a/Mathlib/RingTheory/Localization/Basic.lean b/Mathlib/RingTheory/Localization/Basic.lean index c7eb89c531361..1e220337d2b73 100644 --- a/Mathlib/RingTheory/Localization/Basic.lean +++ b/Mathlib/RingTheory/Localization/Basic.lean @@ -979,7 +979,7 @@ def mkAddMonoidHom (b : M) : R →+ Localization M where theorem mk_sum {ι : Type*} (f : ι → R) (s : Finset ι) (b : M) : mk (∑ i in s, f i) b = ∑ i in s, mk (f i) b := - (mkAddMonoidHom b).map_sum f s + map_sum (mkAddMonoidHom b) f s #align localization.mk_sum Localization.mk_sum theorem mk_list_sum (l : List R) (b : M) : mk l.sum b = (l.map fun a => mk a b).sum := diff --git a/Mathlib/RingTheory/Localization/Integer.lean b/Mathlib/RingTheory/Localization/Integer.lean index 7567b90ba1cb1..ec4ca8ed6e21d 100644 --- a/Mathlib/RingTheory/Localization/Integer.lean +++ b/Mathlib/RingTheory/Localization/Integer.lean @@ -98,7 +98,7 @@ theorem exist_integer_multiples {ι : Type*} (s : Finset ι) (f : ι → S) : · exact (∏ j in s.erase i, (sec M (f j)).2) * (sec M (f i)).1 rw [RingHom.map_mul, sec_spec', ← mul_assoc, ← (algebraMap R S).map_mul, ← Algebra.smul_def] congr 2 - refine' _root_.trans _ ((Submonoid.subtype M).map_prod _ _).symm + refine' _root_.trans _ (map_prod (Submonoid.subtype M) _ _).symm rw [mul_comm,Submonoid.coe_finset_prod, -- Porting note: explicitly supplied `f` ← Finset.prod_insert (f := fun i => ((sec M (f i)).snd : R)) (s.not_mem_erase i), diff --git a/Mathlib/RingTheory/Multiplicity.lean b/Mathlib/RingTheory/Multiplicity.lean index 4fd198ee87b63..e3d31fe977d63 100644 --- a/Mathlib/RingTheory/Multiplicity.lean +++ b/Mathlib/RingTheory/Multiplicity.lean @@ -4,8 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Robert Y. Lewis, Chris Hughes -/ import Mathlib.Algebra.Associated -import Mathlib.Algebra.BigOperators.Basic -import Mathlib.RingTheory.Valuation.Basic +import Mathlib.Algebra.SMulWithZero +import Mathlib.Data.Nat.PartENat +import Mathlib.Tactic.Linarith #align_import ring_theory.multiplicity from "leanprover-community/mathlib"@"e8638a0fcaf73e4500469f368ef9494e495099b3" @@ -440,7 +441,7 @@ theorem multiplicity_add_of_gt {p a b : α} (h : multiplicity p b < multiplicity rw_mod_cast [multiplicity_lt_iff_neg_dvd, dvd_add_right] intro h_dvd · apply multiplicity.is_greatest _ h_dvd - rw [hk, ←Nat.succ_eq_add_one] + rw [hk, ← Nat.succ_eq_add_one] norm_cast apply Nat.lt_succ_self k · rw [pow_dvd_iff_le_multiplicity, Nat.cast_add, ← hk, Nat.cast_one] @@ -623,23 +624,6 @@ theorem multiplicity_pow_self_of_prime {p : α} (hp : Prime p) (n : ℕ) : end CancelCommMonoidWithZero -section Valuation - -variable {R : Type*} [CommRing R] [IsDomain R] {p : R} [DecidableRel (Dvd.dvd : R → R → Prop)] - -/-- `multiplicity` of a prime in an integral domain as an additive valuation to `PartENat`. -/ -noncomputable def addValuation (hp : Prime p) : AddValuation R PartENat := - AddValuation.of (multiplicity p) (multiplicity.zero _) (one_right hp.not_unit) - (fun _ _ => min_le_multiplicity_add) fun _ _ => multiplicity.mul hp -#align multiplicity.add_valuation multiplicity.addValuation - -@[simp] -theorem addValuation_apply {hp : Prime p} {r : R} : addValuation hp r = multiplicity p r := - rfl -#align multiplicity.add_valuation_apply multiplicity.addValuation_apply - -end Valuation - end multiplicity section Nat diff --git a/Mathlib/RingTheory/Nilpotent.lean b/Mathlib/RingTheory/Nilpotent.lean index 8d9379f553ae4..e3e20b632a252 100644 --- a/Mathlib/RingTheory/Nilpotent.lean +++ b/Mathlib/RingTheory/Nilpotent.lean @@ -62,7 +62,7 @@ lemma IsNilpotent.pow {n : ℕ} {S : Type*} [MonoidWithZero S] {x : S} (hx : IsNilpotent x) : IsNilpotent (x ^ n.succ) := by obtain ⟨N,hN⟩ := hx use N - rw [←pow_mul, Nat.succ_mul, pow_add, hN, mul_zero] + rw [← pow_mul, Nat.succ_mul, pow_add, hN, mul_zero] lemma IsNilpotent.pow_of_pos {n} {S : Type*} [MonoidWithZero S] {x : S} (hx : IsNilpotent x) (hn : n ≠ 0) : IsNilpotent (x ^ n) := by @@ -205,7 +205,9 @@ protected lemma isNilpotent_sum {ι : Type*} {s : Finset ι} {f : ι → R} (hnp : ∀ i ∈ s, IsNilpotent (f i)) (h_comm : ∀ i j, i ∈ s → j ∈ s → Commute (f i) (f j)) : IsNilpotent (∑ i in s, f i) := by classical - induction' s using Finset.induction with j s hj ih; simp + induction s using Finset.induction with + | empty => simp + | @insert j s hj ih => ?_ rw [Finset.sum_insert hj] apply Commute.isNilpotent_add · exact Commute.sum_right _ _ _ (fun i hi ↦ h_comm _ _ (by simp) (by simp [hi])) diff --git a/Mathlib/RingTheory/Noetherian.lean b/Mathlib/RingTheory/Noetherian.lean index 86dbede59d0f1..2a05a4fba7c95 100644 --- a/Mathlib/RingTheory/Noetherian.lean +++ b/Mathlib/RingTheory/Noetherian.lean @@ -410,7 +410,7 @@ theorem finite_of_linearIndependent [Nontrivial R] {s : Set M} exact ⟨fun hab x (hxa : x ≤ a) => le_trans hxa hab, fun hx => hx a (le_refl a)⟩ exact ⟨⟨fun n => span R (coe' ∘ f '' { m | m ≤ n }), fun x y => by - rw [le_antisymm_iff, (this x y).symm, (this y x).symm, ←le_antisymm_iff, imp_self] + rw [le_antisymm_iff, (this x y).symm, (this y x).symm, ← le_antisymm_iff, imp_self] trivial⟩, by dsimp [GT.gt]; simp only [lt_iff_le_not_le, (this _ _).symm]; tauto⟩ #align finite_of_linear_independent finite_of_linearIndependent diff --git a/Mathlib/RingTheory/Polynomial/Content.lean b/Mathlib/RingTheory/Polynomial/Content.lean index d5dd0032000d8..521b186c90a94 100644 --- a/Mathlib/RingTheory/Polynomial/Content.lean +++ b/Mathlib/RingTheory/Polynomial/Content.lean @@ -378,7 +378,7 @@ theorem content_mul {p q : R[X]} : (p * q).content = p.content * q.content := by rcases hpq with (hlt | heq) · apply ih _ _ hlt rw [← p.natDegree_primPart, ← q.natDegree_primPart, ← Nat.cast_inj (R := WithBot ℕ), - Nat.cast_add, ←degree_eq_natDegree p.primPart_ne_zero, + Nat.cast_add, ← degree_eq_natDegree p.primPart_ne_zero, ← degree_eq_natDegree q.primPart_ne_zero] at heq rw [p.eq_C_content_mul_primPart, q.eq_C_content_mul_primPart] suffices h : (q.primPart * p.primPart).content = 1 @@ -436,7 +436,7 @@ theorem exists_primitive_lcm_of_isPrimitive {p q : R[X]} (hp : p.IsPrimitive) (h suffices hs : ∀ (n : ℕ) (s : R[X]), s.natDegree = n → p ∣ s ∧ q ∣ s → r ∣ s · apply hs s.natDegree s rfl clear s - by_contra' con + by_contra! con rcases Nat.find_spec con with ⟨s, sdeg, ⟨ps, qs⟩, rs⟩ have s0 : s ≠ 0 := by contrapose! rs diff --git a/Mathlib/RingTheory/Polynomial/Cyclotomic/Eval.lean b/Mathlib/RingTheory/Polynomial/Cyclotomic/Eval.lean index c50c18f560f90..eeb0c4cc212ab 100644 --- a/Mathlib/RingTheory/Polynomial/Cyclotomic/Eval.lean +++ b/Mathlib/RingTheory/Polynomial/Cyclotomic/Eval.lean @@ -62,7 +62,7 @@ private theorem cyclotomic_neg_one_pos {n : ℕ} (hn : 2 < n) {R} [LinearOrdered simp only [Int.cast_one, Int.cast_neg] have h0 := cyclotomic_coeff_zero ℝ hn.le rw [coeff_zero_eq_eval_zero] at h0 - by_contra' hx + by_contra! hx have := intermediate_value_univ (-1) 0 (cyclotomic n ℝ).continuous obtain ⟨y, hy : IsRoot _ y⟩ := this (show (0 : ℝ) ∈ Set.Icc _ _ by simpa [h0] using hx) rw [@isRoot_cyclotomic_iff] at hy diff --git a/Mathlib/RingTheory/Polynomial/Nilpotent.lean b/Mathlib/RingTheory/Polynomial/Nilpotent.lean index 4ccf8ea0765a1..72c4ee6eb3d5a 100644 --- a/Mathlib/RingTheory/Polynomial/Nilpotent.lean +++ b/Mathlib/RingTheory/Polynomial/Nilpotent.lean @@ -72,11 +72,13 @@ protected lemma isNilpotent_iff : replace hp : eval 0 p = 0 := by rwa [coeff_zero_eq_aeval_zero] at hp₀ refine' isNilpotent_C_iff.mpr ⟨k, _⟩ simpa [coeff_zero_eq_aeval_zero, hp] using congr_arg (fun q ↦ coeff q 0) hk - cases' i with i; simpa [hp₀] using hr + cases' i with i + · simpa [hp₀] using hr simp only [coeff_add, coeff_C_succ, add_zero] apply hp simpa using Commute.isNilpotent_sub (Commute.all _ _) hpr hr - · cases' i with i; simp + · cases' i with i + · simp simpa using hnp (isNilpotent_mul_X_iff.mp hpX) i @[simp] lemma isNilpotent_reflect_iff {P : R[X]} {N : ℕ} (hN : P.natDegree ≤ N): @@ -159,8 +161,8 @@ theorem isUnit_iff_coeff_isUnit_isNilpotent : lemma isUnit_iff' : IsUnit P ↔ IsUnit (eval 0 P) ∧ IsNilpotent (P /ₘ X) := by - suffices : P = C (eval 0 P) + X * (P /ₘ X) - · conv_lhs => rw [this]; simp + suffices P = C (eval 0 P) + X * (P /ₘ X) by + conv_lhs => rw [this]; simp conv_lhs => rw [← modByMonic_add_div P monic_X] simp [modByMonic_X] diff --git a/Mathlib/RingTheory/Polynomial/Vieta.lean b/Mathlib/RingTheory/Polynomial/Vieta.lean index fc49834d5d851..80d503ffc882b 100644 --- a/Mathlib/RingTheory/Polynomial/Vieta.lean +++ b/Mathlib/RingTheory/Polynomial/Vieta.lean @@ -171,7 +171,7 @@ theorem MvPolynomial.prod_C_add_X_eq_sum_esymm : (MvPolynomial.esymm σ R j) * Polynomial.X ^ (card σ - j) := by let s := Finset.univ.val.map fun i : σ => (MvPolynomial.X i : MvPolynomial σ R) have : Fintype.card σ = Multiset.card s := by - rw [Multiset.card_map, ←Finset.card_univ, Finset.card_def] + rw [Multiset.card_map, ← Finset.card_univ, Finset.card_def] simp_rw [this, MvPolynomial.esymm_eq_multiset_esymm σ R, Finset.prod_eq_multiset_prod] convert Multiset.prod_X_add_C_eq_sum_esymm s simp_rw [Multiset.map_map, Function.comp_apply] @@ -183,7 +183,7 @@ theorem MvPolynomial.prod_X_add_C_coeff (k : ℕ) (h : k ≤ card σ) : MvPolynomial.esymm σ R (card σ - k) := by let s := Finset.univ.val.map fun i => (MvPolynomial.X i : MvPolynomial σ R) have : Fintype.card σ = Multiset.card s := by - rw [Multiset.card_map, ←Finset.card_univ, Finset.card_def] + rw [Multiset.card_map, ← Finset.card_univ, Finset.card_def] rw [this] at h ⊢ rw [MvPolynomial.esymm_eq_multiset_esymm σ R, Finset.prod_eq_multiset_prod] convert Multiset.prod_X_add_C_coeff s h diff --git a/Mathlib/RingTheory/PolynomialAlgebra.lean b/Mathlib/RingTheory/PolynomialAlgebra.lean index 26b4451f000eb..8a8e6b3f0c95e 100644 --- a/Mathlib/RingTheory/PolynomialAlgebra.lean +++ b/Mathlib/RingTheory/PolynomialAlgebra.lean @@ -226,21 +226,21 @@ noncomputable def matPolyEquiv : Matrix n n R[X] ≃ₐ[R] (Matrix n n R)[X] := #align mat_poly_equiv matPolyEquiv @[simp] theorem matPolyEquiv_symm_C (M : Matrix n n R) : matPolyEquiv.symm (C M) = M.map C := by - simp [matPolyEquiv, ←C_eq_algebraMap] + simp [matPolyEquiv, ← C_eq_algebraMap] @[simp] theorem matPolyEquiv_map_C (M : Matrix n n R) : matPolyEquiv (M.map C) = C M := by - rw [←matPolyEquiv_symm_C, AlgEquiv.apply_symm_apply] + rw [← matPolyEquiv_symm_C, AlgEquiv.apply_symm_apply] @[simp] theorem matPolyEquiv_symm_X : matPolyEquiv.symm X = diagonal fun _ : n => (X : R[X]) := by suffices (Matrix.map 1 fun x ↦ X * algebraMap R R[X] x) = diagonal fun _ : n => (X : R[X]) by simpa [matPolyEquiv] - rw [←Matrix.diagonal_one] + rw [← Matrix.diagonal_one] simp [-Matrix.diagonal_one] @[simp] theorem matPolyEquiv_diagonal_X : matPolyEquiv (diagonal fun _ : n => (X : R[X])) = X := by - rw [←matPolyEquiv_symm_X, AlgEquiv.apply_symm_apply] + rw [← matPolyEquiv_symm_X, AlgEquiv.apply_symm_apply] open Finset diff --git a/Mathlib/RingTheory/PowerSeries/Basic.lean b/Mathlib/RingTheory/PowerSeries/Basic.lean index a5b551b1371f2..b37b2f078bd9e 100644 --- a/Mathlib/RingTheory/PowerSeries/Basic.lean +++ b/Mathlib/RingTheory/PowerSeries/Basic.lean @@ -1901,7 +1901,7 @@ theorem natDegree_trunc_lt (f : R⟦X⟧) (n) : (trunc (n + 1) f).natDegree < n intros rw [coeff_trunc] split_ifs with h - · rw [lt_succ, ←not_lt] at h + · rw [lt_succ, ← not_lt] at h contradiction · rfl @@ -1912,7 +1912,7 @@ theorem degree_trunc_lt (f : R⟦X⟧) (n) : (trunc n f).degree < n := by intros rw [coeff_trunc] split_ifs with h - · rw [←not_le] at h + · rw [← not_le] at h contradiction · rfl @@ -2760,8 +2760,8 @@ theorem coe_pow (n : ℕ) : ((φ ^ n : R[X]) : PowerSeries R) = (φ : PowerSerie #align polynomial.coe_pow Polynomial.coe_pow theorem eval₂_C_X_eq_coe : φ.eval₂ (PowerSeries.C R) PowerSeries.X = ↑φ := by - nth_rw 2 [←eval₂_C_X (p := φ)] - rw [←coeToPowerSeries.ringHom_apply, eval₂_eq_sum_range, eval₂_eq_sum_range, map_sum] + nth_rw 2 [← eval₂_C_X (p := φ)] + rw [← coeToPowerSeries.ringHom_apply, eval₂_eq_sum_range, eval₂_eq_sum_range, map_sum] apply Finset.sum_congr rfl intros rw [map_mul, map_pow, coeToPowerSeries.ringHom_apply, @@ -2864,10 +2864,10 @@ theorem trunc_trunc_mul_trunc {n} (f g : R⟦X⟧) : | zero => rw [pow_zero, pow_zero] | succ a ih => - rw [pow_succ, pow_succ, trunc_trunc_mul, ←trunc_trunc_mul_trunc, ih, trunc_trunc_mul_trunc] + rw [pow_succ, pow_succ, trunc_trunc_mul, ← trunc_trunc_mul_trunc, ih, trunc_trunc_mul_trunc] theorem trunc_coe_eq_self {n} {f : R[X]} (hn : natDegree f < n) : trunc n (f : R⟦X⟧) = f := by - rw [←Polynomial.coe_inj] + rw [← Polynomial.coe_inj] ext m rw [coeff_coe, coeff_trunc] split @@ -2888,7 +2888,7 @@ from the truncations of `f` and `g`.-/ theorem coeff_mul_eq_coeff_trunc_mul_trunc₂ {n a b} (f g) (ha : n < a) (hb : n < b) : coeff R n (f * g) = coeff R n (trunc a f * trunc b g) := by symm - rw [←coeff_coe_trunc_of_lt n.lt_succ_self, ←trunc_trunc_mul_trunc, trunc_trunc_of_le f ha, + rw [← coeff_coe_trunc_of_lt n.lt_succ_self, ← trunc_trunc_mul_trunc, trunc_trunc_of_le f ha, trunc_trunc_of_le g hb, trunc_trunc_mul_trunc, coeff_coe_trunc_of_lt n.lt_succ_self] theorem coeff_mul_eq_coeff_trunc_mul_trunc {d n} (f g) (h : d < n) : diff --git a/Mathlib/RingTheory/PowerSeries/Derivative.lean b/Mathlib/RingTheory/PowerSeries/Derivative.lean index a2d3a90e96512..38755454780d5 100644 --- a/Mathlib/RingTheory/PowerSeries/Derivative.lean +++ b/Mathlib/RingTheory/PowerSeries/Derivative.lean @@ -68,8 +68,8 @@ theorem trunc_derivativeFun (f : R⟦X⟧) (n : ℕ) : --A special case of `derivativeFun_mul`, used in its proof. private theorem derivativeFun_coe_mul_coe (f g : R[X]) : derivativeFun (f * g : R⟦X⟧) = f * derivative g + g * derivative f := by - rw [←coe_mul, derivativeFun_coe, derivative_mul, - add_comm, mul_comm _ g, ←coe_mul, ←coe_mul, Polynomial.coe_add] + rw [← coe_mul, derivativeFun_coe, derivative_mul, + add_comm, mul_comm _ g, ← coe_mul, ← coe_mul, Polynomial.coe_add] /-- Leibniz rule for formal power series.-/ theorem derivativeFun_mul (f g : R⟦X⟧) : @@ -80,10 +80,10 @@ theorem derivativeFun_mul (f g : R⟦X⟧) : rw [coeff_derivativeFun, map_add, coeff_mul_eq_coeff_trunc_mul_trunc _ _ (lt_succ_self _), smul_eq_mul, smul_eq_mul, coeff_mul_eq_coeff_trunc_mul_trunc₂ g f.derivativeFun h₂ h₁, coeff_mul_eq_coeff_trunc_mul_trunc₂ f g.derivativeFun h₂ h₁, trunc_derivativeFun, - trunc_derivativeFun, ←map_add, ←derivativeFun_coe_mul_coe, coeff_derivativeFun] + trunc_derivativeFun, ← map_add, ← derivativeFun_coe_mul_coe, coeff_derivativeFun] theorem derivativeFun_one : derivativeFun (1 : R⟦X⟧) = 0 := by - rw [←map_one (C R), derivativeFun_C (1 : R)] + rw [← map_one (C R), derivativeFun_C (1 : R)] theorem derivativeFun_smul (r : R) (f : R⟦X⟧) : derivativeFun (r • f) = r • derivativeFun f := by rw [smul_eq_C_mul, smul_eq_C_mul, derivativeFun_mul, derivativeFun_C, smul_zero, add_zero, @@ -144,8 +144,8 @@ theorem derivative.ext {R} [CommRing R] [NoZeroSMulDivisors ℕ R] {f g} (hD : d rw [coeff_zero_eq_constantCoeff, hc] | succ n => have equ : coeff R n (d⁄dX R f) = coeff R n (d⁄dX R g) := by rw [hD] - rwa [coeff_derivative, coeff_derivative, ←cast_succ, mul_comm, ←nsmul_eq_mul, - mul_comm, ←nsmul_eq_mul, smul_right_inj n.succ_ne_zero] at equ + rwa [coeff_derivative, coeff_derivative, ← cast_succ, mul_comm, ← nsmul_eq_mul, + mul_comm, ← nsmul_eq_mul, smul_right_inj n.succ_ne_zero] at equ @[simp] theorem derivative_inv {R} [CommRing R] (f : R⟦X⟧ˣ) : d⁄dX R ↑f⁻¹ = -(↑f⁻¹ : R⟦X⟧) ^ 2 * d⁄dX R f := by diff --git a/Mathlib/RingTheory/RootsOfUnity/Basic.lean b/Mathlib/RingTheory/RootsOfUnity/Basic.lean index 0af6a97455c71..560304670bb74 100644 --- a/Mathlib/RingTheory/RootsOfUnity/Basic.lean +++ b/Mathlib/RingTheory/RootsOfUnity/Basic.lean @@ -253,7 +253,7 @@ variable {k R} theorem map_rootsOfUnity_eq_pow_self [RingHomClass F R R] (σ : F) (ζ : rootsOfUnity k R) : ∃ m : ℕ, σ (ζ : Rˣ) = ((ζ : Rˣ) : R) ^ m := by obtain ⟨m, hm⟩ := MonoidHom.map_cyclic (restrictRootsOfUnity σ k) - rw [← restrictRootsOfUnity_coe_apply, hm, ←zpow_mod_orderOf, ← Int.toNat_of_nonneg + rw [← restrictRootsOfUnity_coe_apply, hm, ← zpow_mod_orderOf, ← Int.toNat_of_nonneg (m.emod_nonneg (Int.coe_nat_ne_zero.mpr (pos_iff_ne_zero.mp (orderOf_pos ζ)))), zpow_ofNat, rootsOfUnity.coe_pow] exact ⟨(m % orderOf ζ).toNat, rfl⟩ @@ -994,7 +994,8 @@ theorem autToPow_spec (f : S ≃ₐ[R] S) : μ ^ (hμ.autToPow R f : ZMod n).val refine' (_ : ((hμ.toRootsOfUnity : Sˣ) : S) ^ _ = _).trans this.symm rw [← rootsOfUnity.coe_pow, ← rootsOfUnity.coe_pow] congr 2 - rw [pow_eq_pow_iff_modEq, ZMod.val_nat_cast, hμ.eq_orderOf, ←Subgroup.orderOf_coe, ←orderOf_units] + rw [pow_eq_pow_iff_modEq, ZMod.val_nat_cast, hμ.eq_orderOf, ← Subgroup.orderOf_coe, + ← orderOf_units] exact Nat.mod_modEq _ _ #align is_primitive_root.aut_to_pow_spec IsPrimitiveRoot.autToPow_spec diff --git a/Mathlib/RingTheory/Subsemiring/Basic.lean b/Mathlib/RingTheory/Subsemiring/Basic.lean index 90e8bfa7521e3..7fdf2b8ff5373 100644 --- a/Mathlib/RingTheory/Subsemiring/Basic.lean +++ b/Mathlib/RingTheory/Subsemiring/Basic.lean @@ -46,7 +46,7 @@ theorem natCast_mem [AddSubmonoidWithOneClass S R] (n : ℕ) : (n : R) ∈ s := @[aesop safe apply (rule_sets [SetLike])] lemma ofNat_mem [AddSubmonoidWithOneClass S R] (s : S) (n : ℕ) [n.AtLeastTwo] : no_index (OfNat.ofNat n) ∈ s := by - rw [←Nat.cast_eq_ofNat]; exact natCast_mem s n + rw [← Nat.cast_eq_ofNat]; exact natCast_mem s n instance (priority := 74) AddSubmonoidWithOneClass.toAddMonoidWithOne [AddSubmonoidWithOneClass S R] : AddMonoidWithOne s := diff --git a/Mathlib/RingTheory/TensorProduct.lean b/Mathlib/RingTheory/TensorProduct.lean index 133f63a7e3aeb..012abb6ecb3b5 100644 --- a/Mathlib/RingTheory/TensorProduct.lean +++ b/Mathlib/RingTheory/TensorProduct.lean @@ -179,7 +179,7 @@ instance instAddCommMonoidWithOne : AddCommMonoidWithOne (A ⊗[R] B) where theorem natCast_def (n : ℕ) : (n : A ⊗[R] B) = (n : A) ⊗ₜ (1 : B) := rfl theorem natCast_def' (n : ℕ) : (n : A ⊗[R] B) = (1 : A) ⊗ₜ (n : B) := by - rw [natCast_def, ←nsmul_one, smul_tmul, nsmul_one] + rw [natCast_def, ← nsmul_one, smul_tmul, nsmul_one] end AddCommMonoidWithOne @@ -448,7 +448,7 @@ theorem ext ⦃f g : (A ⊗[R] B) →ₐ[S] C⦄ ext a b have := congr_arg₂ HMul.hMul (AlgHom.congr_fun ha a) (AlgHom.congr_fun hb b) dsimp at * - rwa [←f.map_mul, ←g.map_mul, tmul_mul_tmul, _root_.one_mul, _root_.mul_one] at this + rwa [← f.map_mul, ← g.map_mul, tmul_mul_tmul, _root_.one_mul, _root_.mul_one] at this theorem ext' {g h : A ⊗[R] B →ₐ[S] C} (H : ∀ a b, g (a ⊗ₜ b) = h (a ⊗ₜ b)) : g = h := ext (AlgHom.ext <| fun _ => H _ _) (AlgHom.ext <| fun _ => H _ _) @@ -519,7 +519,7 @@ instance instRing : Ring (A ⊗[R] B) where __ := instNonAssocRing theorem intCast_def' (z : ℤ) : (z : A ⊗[R] B) = (1 : A) ⊗ₜ (z : B) := by - rw [intCast_def, ←zsmul_one, smul_tmul, zsmul_one] + rw [intCast_def, ← zsmul_one, smul_tmul, zsmul_one] -- verify there are no diamonds example : (instRing : Ring (A ⊗[R] B)).toAddCommGroup = addCommGroup := rfl diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain.lean b/Mathlib/RingTheory/UniqueFactorizationDomain.lean index b1a68e1280729..80a68372e8d19 100644 --- a/Mathlib/RingTheory/UniqueFactorizationDomain.lean +++ b/Mathlib/RingTheory/UniqueFactorizationDomain.lean @@ -1920,10 +1920,6 @@ section open Associates UniqueFactorizationMonoid -theorem Associates.quot_out {α : Type*} [CommMonoid α] (a : Associates α) : - Associates.mk (Quot.out a) = a := by rw [← quot_mk_eq_mk, Quot.out_eq] -#align associates.quot_out Associates.quot_out - /-- `toGCDMonoid` constructs a GCD monoid out of a unique factorization domain. -/ noncomputable def UniqueFactorizationMonoid.toGCDMonoid (α : Type*) [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] [DecidableEq (Associates α)] [DecidableEq α] : GCDMonoid α where diff --git a/Mathlib/RingTheory/Valuation/Basic.lean b/Mathlib/RingTheory/Valuation/Basic.lean index 5e25e1bb951f9..90eea6f7c4739 100644 --- a/Mathlib/RingTheory/Valuation/Basic.lean +++ b/Mathlib/RingTheory/Valuation/Basic.lean @@ -499,7 +499,7 @@ theorem isEquiv_iff_val_lt_one [LinearOrderedCommGroupWithZero Γ₀] cases ne_iff_lt_or_gt.1 h_1 with | inl h_2 => simpa [hh, lt_self_iff_false] using h.2 h_2 | inr h_2 => - rw [← inv_one, ←inv_eq_iff_eq_inv, ← map_inv₀] at hh + rw [← inv_one, ← inv_eq_iff_eq_inv, ← map_inv₀] at hh exact hh.not_lt (h.2 ((one_lt_val_iff v' hx).1 h_2)) · intro hh by_contra h_1 @@ -564,7 +564,7 @@ instance [Nontrivial Γ₀] [NoZeroDivisors Γ₀] : Ideal.IsPrime (supp v) := one_ne_zero (α := Γ₀) <| calc 1 = v 1 := v.map_one.symm - _ = 0 := by rw [←mem_supp_iff, h]; exact Submodule.mem_top, + _ = 0 := by rw [← mem_supp_iff, h]; exact Submodule.mem_top, fun {x y} hxy => by simp only [mem_supp_iff] at hxy ⊢ rw [v.map_mul x y] at hxy diff --git a/Mathlib/RingTheory/Valuation/PrimeMultiplicity.lean b/Mathlib/RingTheory/Valuation/PrimeMultiplicity.lean new file mode 100644 index 0000000000000..6df3941e4de03 --- /dev/null +++ b/Mathlib/RingTheory/Valuation/PrimeMultiplicity.lean @@ -0,0 +1,25 @@ +/- +Copyright (c) 2018 Robert Y. Lewis. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Robert Y. Lewis, Chris Hughes +-/ +import Mathlib.RingTheory.Multiplicity +import Mathlib.RingTheory.Valuation.Basic + +/-! +# `multiplicity` of a prime in an integral domain as an additive valuation +-/ + +variable {R : Type*} [CommRing R] [IsDomain R] {p : R} [DecidableRel (Dvd.dvd : R → R → Prop)] + +/-- `multiplicity` of a prime in an integral domain as an additive valuation to `PartENat`. -/ +noncomputable def multiplicity.addValuation (hp : Prime p) : AddValuation R PartENat := + AddValuation.of (multiplicity p) (multiplicity.zero _) (one_right hp.not_unit) + (fun _ _ => min_le_multiplicity_add) fun _ _ => multiplicity.mul hp +#align multiplicity.add_valuation multiplicity.addValuation + +@[simp] +theorem multiplicity.addValuation_apply {hp : Prime p} {r : R} : + addValuation hp r = multiplicity p r := + rfl +#align multiplicity.add_valuation_apply multiplicity.addValuation_apply diff --git a/Mathlib/RingTheory/WittVector/Domain.lean b/Mathlib/RingTheory/WittVector/Domain.lean index 4593e1aa9118a..d17531348db65 100644 --- a/Mathlib/RingTheory/WittVector/Domain.lean +++ b/Mathlib/RingTheory/WittVector/Domain.lean @@ -88,7 +88,7 @@ theorem eq_iterate_verschiebung {x : 𝕎 R} {n : ℕ} (h : ∀ i < n, x.coeff i theorem verschiebung_nonzero {x : 𝕎 R} (hx : x ≠ 0) : ∃ n : ℕ, ∃ x' : 𝕎 R, x'.coeff 0 ≠ 0 ∧ x = verschiebung^[n] x' := by have hex : ∃ k : ℕ, x.coeff k ≠ 0 := by - by_contra' hall + by_contra! hall apply hx ext i simp only [hall, zero_coeff] diff --git a/Mathlib/SetTheory/Cardinal/Basic.lean b/Mathlib/SetTheory/Cardinal/Basic.lean index ba1e73a6190cd..8145fddfb5a20 100644 --- a/Mathlib/SetTheory/Cardinal/Basic.lean +++ b/Mathlib/SetTheory/Cardinal/Basic.lean @@ -629,11 +629,11 @@ theorem lift_bit1 (a : Cardinal) : lift.{v} (bit1 a) = bit1 (lift.{v} a) := by s end deprecated -- Porting note: Proof used to be simp, needed to remind simp that 1 + 1 = 2 -theorem lift_two : lift.{u, v} 2 = 2 := by simp [←one_add_one_eq_two] +theorem lift_two : lift.{u, v} 2 = 2 := by simp [← one_add_one_eq_two] #align cardinal.lift_two Cardinal.lift_two @[simp] -theorem mk_set {α : Type u} : #(Set α) = 2 ^ #α := by simp [←one_add_one_eq_two, Set, mk_arrow] +theorem mk_set {α : Type u} : #(Set α) = 2 ^ #α := by simp [← one_add_one_eq_two, Set, mk_arrow] #align cardinal.mk_set Cardinal.mk_set /-- A variant of `Cardinal.mk_set` expressed in terms of a `Set` instead of a `Type`. -/ @@ -643,7 +643,7 @@ theorem mk_powerset {α : Type u} (s : Set α) : #(↥(𝒫 s)) = 2 ^ #(↥s) := #align cardinal.mk_powerset Cardinal.mk_powerset theorem lift_two_power (a : Cardinal) : lift.{v} (2 ^ a) = 2 ^ lift.{v} a := by - simp [←one_add_one_eq_two] + simp [← one_add_one_eq_two] #align cardinal.lift_two_power Cardinal.lift_two_power section OrderProperties @@ -1103,7 +1103,7 @@ theorem lift_sInf (s : Set Cardinal) : lift.{u,v} (sInf s) = sInf (lift.{u,v} '' theorem lift_iInf {ι} (f : ι → Cardinal) : lift.{u,v} (iInf f) = ⨅ i, lift.{u,v} (f i) := by unfold iInf convert lift_sInf (range f) - simp_rw [←comp_apply (f := lift), range_comp] + simp_rw [← comp_apply (f := lift), range_comp] #align cardinal.lift_infi Cardinal.lift_iInf theorem lift_down {a : Cardinal.{u}} {b : Cardinal.{max u v}} : @@ -1449,7 +1449,7 @@ theorem isLimit_aleph0 : IsLimit ℵ₀ := #align cardinal.is_limit_aleph_0 Cardinal.isLimit_aleph0 theorem IsLimit.aleph0_le {c : Cardinal} (h : IsLimit c) : ℵ₀ ≤ c := by - by_contra' h' + by_contra! h' rcases lt_aleph0.1 h' with ⟨_ | n, rfl⟩ · exact h.ne_zero.irrefl · rw [nat_succ] at h @@ -1703,7 +1703,7 @@ def toNat : ZeroHom Cardinal ℕ where lemma toNat_eq_zero : toNat c = 0 ↔ c = 0 ∨ ℵ₀ ≤ c := by simp only [toNat, ZeroHom.coe_mk, dite_eq_right_iff, or_iff_not_imp_right, not_le] refine' forall_congr' fun h => _ - rw [←@Nat.cast_eq_zero Cardinal, ← Classical.choose_spec (p := fun n : ℕ ↦ c = n)] + rw [← @Nat.cast_eq_zero Cardinal, ← Classical.choose_spec (p := fun n : ℕ ↦ c = n)] lemma toNat_ne_zero : toNat c ≠ 0 ↔ c ≠ 0 ∧ c < ℵ₀ := by simp [not_or] @[simp] lemma toNat_pos : 0 < toNat c ↔ c ≠ 0 ∧ c < ℵ₀ := pos_iff_ne_zero.trans toNat_ne_zero @@ -1757,8 +1757,10 @@ theorem toNat_cast (n : ℕ) : Cardinal.toNat n = n := by exact (Classical.choose_spec (lt_aleph0.1 (nat_lt_aleph0 n))).symm #align cardinal.to_nat_cast Cardinal.toNat_cast +-- See note [no_index around OfNat.ofNat] @[simp] -theorem toNat_ofNat (n : ℕ) [n.AtLeastTwo] : Cardinal.toNat (OfNat.ofNat n) = OfNat.ofNat n := +theorem toNat_ofNat (n : ℕ) [n.AtLeastTwo] : + Cardinal.toNat (no_index (OfNat.ofNat n)) = OfNat.ofNat n := toNat_cast n /-- `toNat` has a right-inverse: coercion. -/ diff --git a/Mathlib/SetTheory/Cardinal/Cofinality.lean b/Mathlib/SetTheory/Cardinal/Cofinality.lean index 7104e6d954a95..5ad120c50dc9c 100644 --- a/Mathlib/SetTheory/Cardinal/Cofinality.lean +++ b/Mathlib/SetTheory/Cardinal/Cofinality.lean @@ -640,7 +640,7 @@ theorem exists_fundamental_sequence (a : Ordinal.{u}) : · push_neg at h cases' wo.wf.min_mem _ h with hji hij refine' ⟨typein r' ⟨_, fun k hkj => lt_of_lt_of_le _ hij⟩, typein_lt_type _ _, _⟩ - · by_contra' H + · by_contra! H exact (wo.wf.not_lt_min _ h ⟨IsTrans.trans _ _ _ hkj hji, H⟩) hkj · rwa [bfamilyOfFamily'_typein] #align ordinal.exists_fundamental_sequence Ordinal.exists_fundamental_sequence @@ -781,7 +781,7 @@ theorem cof_univ : cof univ.{u, v} = Cardinal.univ.{u, v} := then s has an unbounded member -/ theorem unbounded_of_unbounded_sUnion (r : α → α → Prop) [wo : IsWellOrder α r] {s : Set (Set α)} (h₁ : Unbounded r <| ⋃₀ s) (h₂ : #s < StrictOrder.cof r) : ∃ x ∈ s, Unbounded r x := by - by_contra' h + by_contra! h simp_rw [not_unbounded_iff] at h let f : s → α := fun x : s => wo.wf.sup x (h x.1 x.2) refine' h₂.not_le (le_trans (csInf_le' ⟨range f, fun x => _, rfl⟩) mk_range_le) @@ -803,7 +803,7 @@ theorem unbounded_of_unbounded_iUnion {α β : Type u} (r : α → α → Prop) theorem infinite_pigeonhole {β α : Type u} (f : β → α) (h₁ : ℵ₀ ≤ #β) (h₂ : #α < (#β).ord.cof) : ∃ a : α, #(f ⁻¹' {a}) = #β := by have : ∃ a, #β ≤ #(f ⁻¹' {a}) := by - by_contra' h + by_contra! h apply mk_univ.not_lt rw [← preimage_univ, ← iUnion_of_singleton, preimage_iUnion] exact diff --git a/Mathlib/SetTheory/Cardinal/Divisibility.lean b/Mathlib/SetTheory/Cardinal/Divisibility.lean index b5199f95c12db..6b8c0748c6b1b 100644 --- a/Mathlib/SetTheory/Cardinal/Divisibility.lean +++ b/Mathlib/SetTheory/Cardinal/Divisibility.lean @@ -151,7 +151,7 @@ theorem isPrimePow_iff {a : Cardinal} : IsPrimePow a ↔ ℵ₀ ≤ a ∨ ∃ n ⟨p, k, nat_is_prime_iff.2 hp, hk, by rw [han]; exact mod_cast h⟩⟩ rintro ⟨p, k, hp, hk, hpk⟩ have key : p ^ (1 : Cardinal) ≤ ↑a := by - rw [←hpk]; apply power_le_power_left hp.ne_zero; exact mod_cast hk + rw [← hpk]; apply power_le_power_left hp.ne_zero; exact mod_cast hk rw [power_one] at key lift p to ℕ using key.trans_lt (nat_lt_aleph0 a) exact ⟨a, rfl, p, k, nat_is_prime_iff.mp hp, hk, mod_cast hpk⟩ diff --git a/Mathlib/SetTheory/Ordinal/Arithmetic.lean b/Mathlib/SetTheory/Ordinal/Arithmetic.lean index f8708e3d7f39f..38b8fbaf35483 100644 --- a/Mathlib/SetTheory/Ordinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Ordinal/Arithmetic.lean @@ -1282,7 +1282,7 @@ theorem ne_sup_iff_lt_sup {ι : Type u} {f : ι → Ordinal.{max u v}} : theorem sup_not_succ_of_ne_sup {ι : Type u} {f : ι → Ordinal.{max u v}} (hf : ∀ i, f i ≠ sup.{_, v} f) {a} (hao : a < sup.{_, v} f) : succ a < sup.{_, v} f := by - by_contra' hoa + by_contra! hoa exact hao.not_le (sup_le fun i => le_of_lt_succ <| (lt_of_le_of_ne (le_sup _ _) (hf i)).trans_le hoa) #align ordinal.sup_not_succ_of_ne_sup Ordinal.sup_not_succ_of_ne_sup @@ -1623,7 +1623,7 @@ theorem sup_eq_lsub_or_sup_succ_eq_lsub {ι : Type u} (f : ι → Ordinal.{max u theorem sup_succ_le_lsub {ι : Type u} (f : ι → Ordinal.{max u v}) : succ (sup.{_, v} f) ≤ lsub.{_, v} f ↔ ∃ i, f i = sup.{_, v} f := by refine' ⟨fun h => _, _⟩ - · by_contra' hf + · by_contra! hf exact (succ_le_iff.1 h).ne ((sup_le_lsub f).antisymm (lsub_le (ne_sup_iff_lt_sup.1 hf))) rintro ⟨_, hf⟩ rw [succ_le_iff, ← hf] @@ -1640,7 +1640,7 @@ theorem sup_eq_lsub_iff_succ {ι : Type u} (f : ι → Ordinal.{max u v}) : refine' ⟨fun h => _, fun hf => le_antisymm (sup_le_lsub f) (lsub_le fun i => _)⟩ · rw [← h] exact fun a => sup_not_succ_of_ne_sup fun i => (lsub_le_iff.1 (le_of_eq h.symm) i).ne - by_contra' hle + by_contra! hle have heq := (sup_succ_eq_lsub f).2 ⟨i, le_antisymm (le_sup _ _) hle⟩ have := hf _ @@ -1717,7 +1717,7 @@ theorem nonempty_compl_range {ι : Type u} (f : ι → Ordinal.{max u v}) : (Set theorem lsub_typein (o : Ordinal) : lsub.{u, u} (typein ((· < ·) : o.out.α → o.out.α → Prop)) = o := (lsub_le.{u, u} typein_lt_self).antisymm (by - by_contra' h + by_contra! h -- Porting note: `nth_rw` → `conv_rhs` & `rw` conv_rhs at h => rw [← type_lt o] simpa [typein_enum] using lt_lsub.{u, u} (typein (· < ·)) (enum (· < ·) _ h)) @@ -1839,7 +1839,7 @@ theorem bsup_eq_blsub_or_succ_bsup_eq_blsub {o : Ordinal.{u}} (f : ∀ a < o, Or theorem bsup_succ_le_blsub {o : Ordinal.{u}} (f : ∀ a < o, Ordinal.{max u v}) : succ (bsup.{_, v} o f) ≤ blsub.{_, v} o f ↔ ∃ i hi, f i hi = bsup.{_, v} o f := by refine' ⟨fun h => _, _⟩ - · by_contra' hf + · by_contra! hf exact ne_of_lt (succ_le_iff.1 h) (le_antisymm (bsup_le_blsub f) (blsub_le (lt_bsup_of_ne_bsup.1 hf))) @@ -2027,7 +2027,7 @@ theorem mex_not_mem_range {ι : Type u} (f : ι → Ordinal.{max u v}) : mex.{_, theorem le_mex_of_forall {ι : Type u} {f : ι → Ordinal.{max u v}} {a : Ordinal} (H : ∀ b < a, ∃ i, f i = b) : a ≤ mex.{_, v} f := by - by_contra' h + by_contra! h exact mex_not_mem_range f (H _ h) #align ordinal.le_mex_of_forall Ordinal.le_mex_of_forall @@ -2040,7 +2040,7 @@ theorem mex_le_of_ne {ι} {f : ι → Ordinal} {a} (ha : ∀ i, f i ≠ a) : mex #align ordinal.mex_le_of_ne Ordinal.mex_le_of_ne theorem exists_of_lt_mex {ι} {f : ι → Ordinal} {a} (ha : a < mex f) : ∃ i, f i = a := by - by_contra' ha' + by_contra! ha' exact ha.not_le (mex_le_of_ne ha') #align ordinal.exists_of_lt_mex Ordinal.exists_of_lt_mex @@ -2058,7 +2058,7 @@ theorem mex_monotone {α β : Type u} {f : α → Ordinal.{max u v}} {g : β → theorem mex_lt_ord_succ_mk {ι : Type u} (f : ι → Ordinal.{u}) : mex.{_, u} f < (succ #ι).ord := by - by_contra' h + by_contra! h apply (lt_succ #ι).not_le have H := fun a => exists_of_lt_mex ((typein_lt_self a).trans_le h) let g : (succ #ι).ord.out.α → ι := fun a => Classical.choose (H a) @@ -2088,7 +2088,7 @@ theorem bmex_not_mem_brange {o : Ordinal} (f : ∀ a < o, Ordinal) : bmex o f theorem le_bmex_of_forall {o : Ordinal} (f : ∀ a < o, Ordinal) {a : Ordinal} (H : ∀ b < a, ∃ i hi, f i hi = b) : a ≤ bmex o f := by - by_contra' h + by_contra! h exact bmex_not_mem_brange f (H _ h) #align ordinal.le_bmex_of_forall Ordinal.le_bmex_of_forall @@ -2268,7 +2268,7 @@ theorem enumOrd_surjective (hS : Unbounded (· < ·) S) : ∀ s ∈ S, ∃ a, en -- Porting note: `flip` is required to infer a metavariable. rcases flip exists_lt_of_lt_csSup ha ⟨0, this⟩ with ⟨b, hb, hab⟩ exact (enumOrd_strictMono hS hab).trans_le hb - · by_contra' h + · by_contra! h exact (le_csSup ⟨s, fun a => (lt_wf.self_le_of_strictMono (enumOrd_strictMono hS) a).trans⟩ (enumOrd_succ_le hS hs h)).not_lt @@ -2514,7 +2514,7 @@ theorem add_le_of_forall_add_lt {a b c : Ordinal} (hb : 0 < b) (h : ∀ d < b, a exact (h _ hb).le) rw [← H] apply add_le_add_left _ a - by_contra' hb + by_contra! hb exact (h _ hb).ne H #align ordinal.add_le_of_forall_add_lt Ordinal.add_le_of_forall_add_lt diff --git a/Mathlib/SetTheory/Ordinal/Basic.lean b/Mathlib/SetTheory/Ordinal/Basic.lean index c548553931af5..10446df87444a 100644 --- a/Mathlib/SetTheory/Ordinal/Basic.lean +++ b/Mathlib/SetTheory/Ordinal/Basic.lean @@ -1509,7 +1509,7 @@ theorem ord_univ : ord univ.{u, v} = Ordinal.univ.{u, v} := by refine' le_antisymm (ord_card_le _) <| le_of_forall_lt fun o h => lt_ord.2 ?_ have := lift.principalSeg.{u, v}.down.1 (by simpa only [lift.principalSeg_coe] using h) rcases this with ⟨o, h'⟩ - rw [←h', lift.principalSeg_coe, ← lift_card] + rw [← h', lift.principalSeg_coe, ← lift_card] apply lift_lt_univ' #align cardinal.ord_univ Cardinal.ord_univ diff --git a/Mathlib/SetTheory/Ordinal/Exponential.lean b/Mathlib/SetTheory/Ordinal/Exponential.lean index 1001103c7694c..6d08cedf7a25c 100644 --- a/Mathlib/SetTheory/Ordinal/Exponential.lean +++ b/Mathlib/SetTheory/Ordinal/Exponential.lean @@ -413,7 +413,7 @@ theorem opow_mul_add_lt_opow_succ {b u v w : Ordinal} (hvb : v < b) (hw : w < b theorem log_opow_mul_add {b u v w : Ordinal} (hb : 1 < b) (hv : v ≠ 0) (hvb : v < b) (hw : w < b ^ u) : log b (b ^ u * v + w) = u := by have hne' := (opow_mul_add_pos (zero_lt_one.trans hb).ne' u hv w).ne' - by_contra' hne + by_contra! hne cases' lt_or_gt_of_ne hne with h h · rw [← lt_opow_iff_log_lt hb hne'] at h exact h.not_le ((le_mul_left _ (Ordinal.pos_iff_ne_zero.2 hv)).trans (le_add_right _ _)) diff --git a/Mathlib/SetTheory/Ordinal/FixedPoint.lean b/Mathlib/SetTheory/Ordinal/FixedPoint.lean index d0a6d35550777..f3b57f413f144 100644 --- a/Mathlib/SetTheory/Ordinal/FixedPoint.lean +++ b/Mathlib/SetTheory/Ordinal/FixedPoint.lean @@ -292,7 +292,7 @@ theorem nfpBFamily_monotone (hf : ∀ i hi, Monotone (f i hi)) : Monotone (nfpBF theorem apply_lt_nfpBFamily (H : ∀ i hi, IsNormal (f i hi)) {a b} (hb : b < nfpBFamily.{u, v} o f a) (i hi) : f i hi b < nfpBFamily.{u, v} o f a := by - rw [←familyOfBFamily_enum o f] + rw [← familyOfBFamily_enum o f] apply apply_lt_nfpFamily (fun _ => H _ _) hb #align ordinal.apply_lt_nfp_bfamily Ordinal.apply_lt_nfpBFamily diff --git a/Mathlib/SetTheory/Ordinal/NaturalOps.lean b/Mathlib/SetTheory/Ordinal/NaturalOps.lean index 26c85fcc48a33..9f4ac5173675b 100644 --- a/Mathlib/SetTheory/Ordinal/NaturalOps.lean +++ b/Mathlib/SetTheory/Ordinal/NaturalOps.lean @@ -373,7 +373,7 @@ instance add_covariantClass_le : CovariantClass NatOrdinal.{u} NatOrdinal.{u} ( instance add_contravariantClass_le : ContravariantClass NatOrdinal.{u} NatOrdinal.{u} (· + ·) (· ≤ ·) := ⟨fun a b c h => by - by_contra' h' + by_contra! h' exact h.not_lt (add_lt_add_left h' a)⟩ #align nat_ordinal.add_contravariant_class_le NatOrdinal.add_contravariantClass_le diff --git a/Mathlib/SetTheory/Ordinal/Notation.lean b/Mathlib/SetTheory/Ordinal/Notation.lean index e0c8a27543763..ef1fc29051039 100644 --- a/Mathlib/SetTheory/Ordinal/Notation.lean +++ b/Mathlib/SetTheory/Ordinal/Notation.lean @@ -986,7 +986,7 @@ theorem repr_opow (o₁ o₂) [NF o₁] [NF o₂] : repr (o₁ ^ o₂) = repr o simp only [(· ^ ·)] simp [Pow.pow, opow, Ordinal.succ_ne_zero] · simpa using nat_cast_lt.2 (Nat.succ_lt_succ <| pos_iff_ne_zero.2 h) - · rw [←Nat.cast_succ, lt_omega] + · rw [← Nat.cast_succ, lt_omega] exact ⟨_, rfl⟩ · haveI := N₁.fst haveI := N₁.snd diff --git a/Mathlib/SetTheory/Ordinal/Principal.lean b/Mathlib/SetTheory/Ordinal/Principal.lean index 53af177bf8aa8..ef848f271cf68 100644 --- a/Mathlib/SetTheory/Ordinal/Principal.lean +++ b/Mathlib/SetTheory/Ordinal/Principal.lean @@ -162,7 +162,7 @@ theorem exists_lt_add_of_not_principal_add {a} (ha : ¬Principal (· + ·) a) : theorem principal_add_iff_add_lt_ne_self {a} : Principal (· + ·) a ↔ ∀ ⦃b c⦄, b < a → c < a → b + c ≠ a := ⟨fun ha b c hb hc => (ha hb hc).ne, fun H => by - by_contra' ha + by_contra! ha rcases exists_lt_add_of_not_principal_add ha with ⟨b, c, hb, hc, rfl⟩ exact (H hb hc).irrefl⟩ #align ordinal.principal_add_iff_add_lt_ne_self Ordinal.principal_add_iff_add_lt_ne_self diff --git a/Mathlib/SetTheory/Ordinal/Topology.lean b/Mathlib/SetTheory/Ordinal/Topology.lean index 54a456ab2d100..fb0415372b704 100644 --- a/Mathlib/SetTheory/Ordinal/Topology.lean +++ b/Mathlib/SetTheory/Ordinal/Topology.lean @@ -229,7 +229,7 @@ theorem enumOrd_isNormal_iff_isClosed (hs : s.Unbounded (· < ·)) : b hb rw [← hb] apply Hs.monotone - by_contra' hba + by_contra! hba apply (Hs (lt_succ b)).not_le rw [hb] exact le_bsup.{u, u} _ _ (ha.2 _ hba) diff --git a/Mathlib/SetTheory/ZFC/Basic.lean b/Mathlib/SetTheory/ZFC/Basic.lean index fd6cce188eb2f..8bfc3b9f71fad 100644 --- a/Mathlib/SetTheory/ZFC/Basic.lean +++ b/Mathlib/SetTheory/ZFC/Basic.lean @@ -1699,7 +1699,7 @@ theorem sInter_empty : ⋂₀ (∅ : Class.{u}) = univ := by theorem eq_univ_of_powerset_subset {A : Class} (hA : powerset A ⊆ A) : A = univ := eq_univ_of_forall (by - by_contra' hnA + by_contra! hnA exact WellFounded.min_mem ZFSet.mem_wf _ hnA (hA fun x hx => @@ -1786,10 +1786,10 @@ theorem choice_mem (y : ZFSet.{u}) (yx : y ∈ x) : (choice x ′ y : Class.{u}) private lemma toSet_equiv_aux {s : Set ZFSet.{u}} (hs : Small.{u} s) : (mk $ PSet.mk (Shrink s) fun x ↦ ((equivShrink s).symm x).1.out).toSet = s := by ext x - rw [mem_toSet, ←mk_out x, mk_mem_iff, mk_out] + rw [mem_toSet, ← mk_out x, mk_mem_iff, mk_out] refine' ⟨_, λ xs ↦ ⟨equivShrink s (Subtype.mk x xs), _⟩⟩ · rintro ⟨b, h2⟩ - rw [←ZFSet.eq, ZFSet.mk_out] at h2 + rw [← ZFSet.eq, ZFSet.mk_out] at h2 simp [h2] · simp [PSet.Equiv.refl] diff --git a/Mathlib/Tactic/ApplyFun.lean b/Mathlib/Tactic/ApplyFun.lean index 3aaf64ae18b67..09d1d12840f90 100644 --- a/Mathlib/Tactic/ApplyFun.lean +++ b/Mathlib/Tactic/ApplyFun.lean @@ -105,7 +105,7 @@ def maybeProveInjective (ginj : Expr) (using? : Option Expr) : MetaM Bool := do let err ← mkHasTypeButIsExpectedMsg (← inferType u) (← inferType ginj) throwError "Using clause {err}" -- Try an assumption - if ←ginj.mvarId!.assumptionCore then + if ← ginj.mvarId!.assumptionCore then return true -- Try using that this is an equivalence -- Note: if `f` is itself a metavariable, this can cause it to become an equivalence; diff --git a/Mathlib/Tactic/ByContra.lean b/Mathlib/Tactic/ByContra.lean index f69d742c79423..ec7bc62e4cbd9 100644 --- a/Mathlib/Tactic/ByContra.lean +++ b/Mathlib/Tactic/ByContra.lean @@ -10,11 +10,11 @@ open Lean Lean.Parser Parser.Tactic Elab Command Elab.Tactic Meta /-- If the target of the main goal is a proposition `p`, -`by_contra'` reduces the goal to proving `False` using the additional hypothesis `this : ¬ p`. -`by_contra' h` can be used to name the hypothesis `h : ¬ p`. +`by_contra!` reduces the goal to proving `False` using the additional hypothesis `this : ¬ p`. +`by_contra! h` can be used to name the hypothesis `h : ¬ p`. The hypothesis `¬ p` will be negation normalized using `push_neg`. For instance, `¬ a < b` will be changed to `b ≤ a`. -`by_contra' h : q` will normalize negations in `¬ p`, normalize negations in `q`, +`by_contra! h : q` will normalize negations in `¬ p`, normalize negations in `q`, and then check that the two normalized forms are equal. The resulting hypothesis is the pre-normalized form, `q`. If the name `h` is not explicitly provided, then `this` will be used as name. @@ -23,22 +23,22 @@ It is a variant on the tactic `by_contra`. Examples: ```lean example : 1 < 2 := by - by_contra' h + by_contra! h -- h : 2 ≤ 1 ⊢ False example : 1 < 2 := by - by_contra' h : ¬ 1 < 2 + by_contra! h : ¬ 1 < 2 -- h : ¬ 1 < 2 ⊢ False ``` -/ -syntax (name := byContra') "by_contra'" (ppSpace colGt binderIdent)? Term.optType : tactic +syntax (name := byContra!) "by_contra!" (ppSpace colGt binderIdent)? Term.optType : tactic macro_rules - | `(tactic| by_contra'%$tk $[_%$under]? $[: $ty]?) => - `(tactic| by_contra' $(mkIdentFrom (under.getD tk) `this (canonical := true)):ident $[: $ty]?) - | `(tactic| by_contra' $e:ident) => `(tactic| (by_contra $e:ident; push_neg at $e:ident)) - | `(tactic| by_contra' $e:ident : $y) => `(tactic| - (by_contra' h; + | `(tactic| by_contra!%$tk $[_%$under]? $[: $ty]?) => + `(tactic| by_contra! $(mkIdentFrom (under.getD tk) `this (canonical := true)):ident $[: $ty]?) + | `(tactic| by_contra! $e:ident) => `(tactic| (by_contra $e:ident; push_neg at $e:ident)) + | `(tactic| by_contra! $e:ident : $y) => `(tactic| + (by_contra! h; -- if the below `exact` call fails then this tactic should fail with the message -- tactic failed: and are not definitionally equal have $e:ident : $y := by { push_neg; exact h }; diff --git a/Mathlib/Tactic/CancelDenoms/Core.lean b/Mathlib/Tactic/CancelDenoms/Core.lean index 86a2a53665ce9..9a8ebbeb22ffc 100644 --- a/Mathlib/Tactic/CancelDenoms/Core.lean +++ b/Mathlib/Tactic/CancelDenoms/Core.lean @@ -65,10 +65,10 @@ theorem neg_subst {α} [Ring α] {n e t : α} (h1 : n * e = t) : n * -e = -t := theorem pow_subst {α} [CommRing α] {n e1 t1 k l : α} {e2 : ℕ} (h1 : n * e1 = t1) (h2 : l * n ^ e2 = k) : k * (e1 ^ e2) = l * t1 ^ e2 := by - rw [←h2, ←h1, mul_pow, mul_assoc] + rw [← h2, ← h1, mul_pow, mul_assoc] theorem inv_subst {α} [Field α] {n k e : α} (h2 : e ≠ 0) (h3 : n * e = k) : - k * (e ⁻¹) = n := by rw [←div_eq_mul_inv, ←h3, mul_div_cancel _ h2] + k * (e ⁻¹) = n := by rw [← div_eq_mul_inv, ← h3, mul_div_cancel _ h2] theorem cancel_factors_lt {α} [LinearOrderedField α] {a b ad bd a' b' gcd : α} (ha : ad * a = a') (hb : bd * b = b') (had : 0 < ad) (hbd : 0 < bd) (hgcd : 0 < gcd) : diff --git a/Mathlib/Tactic/CategoryTheory/Slice.lean b/Mathlib/Tactic/CategoryTheory/Slice.lean index 76dd5b6a332cb..60a77786bac0a 100644 --- a/Mathlib/Tactic/CategoryTheory/Slice.lean +++ b/Mathlib/Tactic/CategoryTheory/Slice.lean @@ -49,7 +49,7 @@ def evalSlice (a b : Nat) : TacticM Unit := do evalTactic (← `(conv| congr)) evalTactic (← `(tactic| rotate_left)) let k ← iterateUntilFailureCount - <| evalTactic (← `(conv| rw [←Category.assoc])) + <| evalTactic (← `(conv| rw [← Category.assoc])) let c := k+1+a-b iterateRange c c <| evalTactic (← `(conv| congr)) let _ ← iterateUntilFailureWithResults do diff --git a/Mathlib/Tactic/Common.lean b/Mathlib/Tactic/Common.lean index d3bc61458fa11..4ddcec5e60015 100644 --- a/Mathlib/Tactic/Common.lean +++ b/Mathlib/Tactic/Common.lean @@ -39,6 +39,9 @@ import Mathlib.Tactic.ExtractGoal import Mathlib.Tactic.ExtractLets import Mathlib.Tactic.FailIfNoProgress import Mathlib.Tactic.Find +-- `gcongr` currently imports `Algebra.Order.Field.Power` and thence `Algebra.CharZero.Lemmas` +-- Hopefully this can be rearranged. +-- import Mathlib.Tactic.GCongr import Mathlib.Tactic.GeneralizeProofs import Mathlib.Tactic.GuardGoalNums import Mathlib.Tactic.GuardHypNums @@ -57,6 +60,8 @@ import Mathlib.Tactic.MkIffOfInductiveProp -- import Mathlib.Tactic.NormNum.Basic import Mathlib.Tactic.NthRewrite import Mathlib.Tactic.Observe +-- `positivity` imports `Data.Nat.Factorial.Basic`, but hopefully this can be rearranged. +-- import Mathlib.Tactic.Positivity import Mathlib.Tactic.ProjectionNotation import Mathlib.Tactic.Propose import Mathlib.Tactic.PushNeg @@ -98,6 +103,9 @@ import Mathlib.Tactic.TypeCheck import Mathlib.Tactic.UnsetOption import Mathlib.Tactic.Use import Mathlib.Tactic.Variable +import Mathlib.Tactic.Widget.Calc +import Mathlib.Tactic.Widget.Congrm +import Mathlib.Tactic.Widget.Conv import Mathlib.Tactic.WLOG import Mathlib.Util.AssertExists import Mathlib.Util.CountHeartbeats diff --git a/Mathlib/Tactic/Explode/Pretty.lean b/Mathlib/Tactic/Explode/Pretty.lean index 6bac0c26aeb2f..44f15e7531e3d 100644 --- a/Mathlib/Tactic/Explode/Pretty.lean +++ b/Mathlib/Tactic/Explode/Pretty.lean @@ -57,7 +57,7 @@ def entriesToMessageData (entries : Entries) : MetaM MessageData := do -- ['1', '2', '3'] let paddedLines ← padRight <| entries.l.map fun entry => m!"{entry.line!}" -- [' ', '1,2', '1 '] - let paddedDeps ← padRight <| entries.l.map fun entry => + let paddedDeps ← padRight <| entries.l.map fun entry => String.intercalate "," <| entry.deps.map (fun dep => (dep.map toString).getD "_") -- ['p ', 'hP ', '∀I '] let paddedThms ← padRight <| entries.l.map (·.thm) diff --git a/Mathlib/Tactic/Find.lean b/Mathlib/Tactic/Find.lean index 3b414a7065afd..14effbeab0dd2 100644 --- a/Mathlib/Tactic/Find.lean +++ b/Mathlib/Tactic/Find.lean @@ -50,7 +50,7 @@ private def isBlackListed (declName : Name) : MetaM Bool := do <||> isMatcher declName initialize findDeclsPerHead : DeclCache (Lean.HashMap HeadIndex (Array Name)) ← - DeclCache.mk "#find: init cache" {} fun _ c headMap ↦ do + DeclCache.mk "#find: init cache" failure {} fun _ c headMap ↦ do if (← isBlackListed c.name) then return headMap -- TODO: this should perhaps use `forallTelescopeReducing` instead, diff --git a/Mathlib/Tactic/Group.lean b/Mathlib/Tactic/Group.lean index 5c9c750e39b5d..41bf69ce34bef 100644 --- a/Mathlib/Tactic/Group.lean +++ b/Mathlib/Tactic/Group.lean @@ -58,12 +58,12 @@ macro_rules | `(tactic| aux_group₁ $[at $location]?) => `(tactic| simp (config := {decide := false, failIfUnchanged := false}) only [commutatorElement_def, mul_one, one_mul, - ←zpow_neg_one, ←zpow_ofNat, ←zpow_mul, + ← zpow_neg_one, ← zpow_ofNat, ← zpow_mul, Int.ofNat_add, Int.ofNat_mul, Int.mul_neg_eq_neg_mul_symm, Int.neg_mul_eq_neg_mul_symm, neg_neg, one_zpow, zpow_zero, zpow_one, mul_zpow_neg_one, - ←mul_assoc, - ←zpow_add, ←zpow_add_one, ←zpow_one_add, zpow_trick, zpow_trick_one, zpow_trick_one', + ← mul_assoc, + ← zpow_add, ← zpow_add_one, ← zpow_one_add, zpow_trick, zpow_trick_one, zpow_trick_one', tsub_self, sub_self, add_neg_self, neg_add_self] $[at $location]?) diff --git a/Mathlib/Tactic/LibrarySearch.lean b/Mathlib/Tactic/LibrarySearch.lean index 2a582884ee98a..5b53531bb0ec7 100644 --- a/Mathlib/Tactic/LibrarySearch.lean +++ b/Mathlib/Tactic/LibrarySearch.lean @@ -8,6 +8,7 @@ import Std.Util.Cache import Mathlib.Tactic.SolveByElim import Std.Data.MLList.Heartbeats import Mathlib.Lean.Name +import Mathlib.Lean.Meta.DiscrTree /-! # Library search @@ -60,9 +61,10 @@ def processLemma (name : Name) (constInfo : ConstantInfo) : let mut r := #[(keys, (name, .none))] match type.getAppFnArgs with | (``Iff, #[lhs, rhs]) => do - return r.push (← DiscrTree.mkPath rhs discrTreeConfig, (name, .mp)) + r := r.push (← DiscrTree.mkPath rhs discrTreeConfig, (name, .mp)) |>.push (← DiscrTree.mkPath lhs discrTreeConfig, (name, .mpr)) - | _ => return r + | _ => pure () + return r.filter (DiscrTree.keysSpecific ·.1) /-- Construct the discrimination tree of all lemmas. -/ def buildDiscrTree : IO (DiscrTreeCache (Name × DeclMod)) := @@ -90,9 +92,9 @@ Retrieve the current current of lemmas. initialize librarySearchLemmas : DiscrTreeCache (Name × DeclMod) ← unsafe do let path ← cachePath if (← path.pathExists) then - let (d, _r) ← unpickle (DiscrTree (Name × DeclMod)) path - -- We can drop the `CompactedRegion` value; we do not plan to free it - DiscrTreeCache.mk "apply?: using cache" processLemma (init := some d) + -- We can drop the `CompactedRegion` value from `unpickle`; we do not plan to free it + let d := (·.1) <$> unpickle (DiscrTree (Name × DeclMod)) path + DiscrTreeCache.mk "apply?: using cache" processLemma (init := d) else buildDiscrTree diff --git a/Mathlib/Tactic/Linarith/Verification.lean b/Mathlib/Tactic/Linarith/Verification.lean index 8575f647c8d99..b4d337703576a 100644 --- a/Mathlib/Tactic/Linarith/Verification.lean +++ b/Mathlib/Tactic/Linarith/Verification.lean @@ -120,7 +120,7 @@ def mkLTZeroProof : List (Expr × ℕ) → MetaM Expr step (c : Ineq) (pf npf : Expr) (coeff : ℕ) : MetaM (Ineq × Expr) := do let (iq, h') ← mkSingleCompZeroOf coeff npf let (nm, niq) := addIneq c iq - return (niq, ←mkAppM nm #[pf, h']) + return (niq, ← mkAppM nm #[pf, h']) /-- If `prf` is a proof of `t R s`, `leftOfIneqProof prf` returns `t`. -/ def leftOfIneqProof (prf : Expr) : MetaM Expr := do diff --git a/Mathlib/Tactic/MkIffOfInductiveProp.lean b/Mathlib/Tactic/MkIffOfInductiveProp.lean index 6457562047a28..4606e4367f25f 100644 --- a/Mathlib/Tactic/MkIffOfInductiveProp.lean +++ b/Mathlib/Tactic/MkIffOfInductiveProp.lean @@ -78,7 +78,7 @@ args.foldrM (λarg i:Expr => do let l := (← inferType t).sortLevel! if arg.occurs i || l != Level.zero then pure (mkApp2 (mkConst `Exists [l] : Expr) t - (updateLambdaBinderInfoD! <| ←mkLambdaFVars #[arg] i)) + (updateLambdaBinderInfoD! <| ← mkLambdaFVars #[arg] i)) else pure <| mkApp2 (mkConst `And [] : Expr) t i) inner diff --git a/Mathlib/Tactic/Nontriviality/Core.lean b/Mathlib/Tactic/Nontriviality/Core.lean index d6d00ee555af5..2e453cd7d3b07 100644 --- a/Mathlib/Tactic/Nontriviality/Core.lean +++ b/Mathlib/Tactic/Nontriviality/Core.lean @@ -70,19 +70,19 @@ a list of additional `simp` lemmas that can be passed to `nontriviality` using t ``` example {R : Type} [OrderedRing R] {a : R} (h : 0 < a) : 0 < a := by - nontriviality -- There is now a `nontrivial R` hypothesis available. + nontriviality -- There is now a `Nontrivial R` hypothesis available. assumption ``` ``` example {R : Type} [CommRing R] {r s : R} : r * s = s * r := by - nontriviality -- There is now a `nontrivial R` hypothesis available. + nontriviality -- There is now a `Nontrivial R` hypothesis available. apply mul_comm ``` ``` example {R : Type} [OrderedRing R] {a : R} (h : 0 < a) : (2 : ℕ) ∣ 4 := by - nontriviality R -- there is now a `nontrivial R` hypothesis available. + nontriviality R -- there is now a `Nontrivial R` hypothesis available. dec_trivial ``` @@ -91,7 +91,7 @@ def myeq {α : Type} (a b : α) : Prop := a = b example {α : Type} (a b : α) (h : a = b) : myeq a b := by success_if_fail nontriviality α -- Fails - nontriviality α using myeq -- There is now a `nontrivial α` hypothesis available + nontriviality α using myeq -- There is now a `Nontrivial α` hypothesis available assumption ``` -/ diff --git a/Mathlib/Tactic/NormNum/Basic.lean b/Mathlib/Tactic/NormNum/Basic.lean index d2ed3b65d55e5..c0e205000fd71 100644 --- a/Mathlib/Tactic/NormNum/Basic.lean +++ b/Mathlib/Tactic/NormNum/Basic.lean @@ -170,7 +170,7 @@ theorem isRat_add {α} [Ring α] {f : α → α → α} {a b : α} {na nb nc : simp only [Int.cast_add, Int.cast_mul, Int.cast_ofNat, ← mul_assoc, add_mul, mul_mul_invOf_self_cancel] at h₁ have h₂ := congr_arg (↑nc * ↑· * (⅟↑da * ⅟↑db * ⅟↑dc : α)) h₂ - simp only [H, mul_mul_invOf_self_cancel', Nat.cast_mul, ←mul_assoc] at h₁ h₂ + simp only [H, mul_mul_invOf_self_cancel', Nat.cast_mul, ← mul_assoc] at h₁ h₂ rw [h₁, h₂, Nat.cast_commute] simp only [mul_mul_invOf_self_cancel, (Nat.cast_commute (α := α) da dc).invOf_left.invOf_right.right_comm, diff --git a/Mathlib/Tactic/NormNum/BigOperators.lean b/Mathlib/Tactic/NormNum/BigOperators.lean index b9fe5d2053b70..66c39458fe52b 100644 --- a/Mathlib/Tactic/NormNum/BigOperators.lean +++ b/Mathlib/Tactic/NormNum/BigOperators.lean @@ -356,7 +356,7 @@ If your finset is not supported, you can add it to the match in `Finset.proveEmp partial def evalFinsetProd : NormNumExt where eval {u β} e := do let .app (.app (.app (.app (.app (.const `Finset.prod [_, v]) β') α) _) s) f ← whnfR e | failure - guard <| ←withNewMCtxDepth <| isDefEq β β' + guard <| ← withNewMCtxDepth <| isDefEq β β' have α : Q(Type v) := α have s : Q(Finset $α) := s have f : Q($α → $β) := f @@ -385,7 +385,7 @@ If your finset is not supported, you can add it to the match in `Finset.proveEmp partial def evalFinsetSum : NormNumExt where eval {u β} e := do let .app (.app (.app (.app (.app (.const `Finset.sum [_, v]) β') α) _) s) f ← whnfR e | failure - guard <| ←withNewMCtxDepth <| isDefEq β β' + guard <| ← withNewMCtxDepth <| isDefEq β β' have α : Q(Type v) := α have s : Q(Finset $α) := s have f : Q($α → $β) := f diff --git a/Mathlib/Tactic/Propose.lean b/Mathlib/Tactic/Propose.lean index 71ef68a56bc4d..b83a8a6427eaa 100644 --- a/Mathlib/Tactic/Propose.lean +++ b/Mathlib/Tactic/Propose.lean @@ -46,7 +46,7 @@ initialize registerTraceClass `Tactic.propose def discrTreeConfig : WhnfCoreConfig := {} initialize proposeLemmas : DeclCache (DiscrTree Name) ← - DeclCache.mk "have?: init cache" {} fun name constInfo lemmas => do + DeclCache.mk "have?: init cache" failure {} fun name constInfo lemmas => do if constInfo.isUnsafe then return lemmas if ← name.isBlackListed then return lemmas withNewMCtxDepth do withReducible do diff --git a/Mathlib/Tactic/Rewrites.lean b/Mathlib/Tactic/Rewrites.lean index 6d5ba9f50a5b5..3042002b3e5dd 100644 --- a/Mathlib/Tactic/Rewrites.lean +++ b/Mathlib/Tactic/Rewrites.lean @@ -21,7 +21,7 @@ import Mathlib.Tactic.SolveByElim `rw?` should not be left in proofs; it is a search tool, like `apply?`. -Suggestions are printed as `rw [h]` or `rw [←h]`. +Suggestions are printed as `rw [h]` or `rw [← h]`. -/ @@ -56,9 +56,7 @@ def backwardWeight := 1 /-- Configuration for `DiscrTree`. -/ def discrTreeConfig : WhnfCoreConfig := {} -/-- We will discard -/ -def keysSpecific (keys : Array DiscrTree.Key) : Bool := - !(keys == #[.star] || keys == #[.const `Eq 3, .star, .star, .star]) +open Lean.Meta.DiscrTree (keysSpecific) /-- Prepare the discrimination tree entries for a lemma. -/ def processLemma (name : Name) (constInfo : ConstantInfo) : @@ -126,9 +124,9 @@ Retrieve the current cache of lemmas. initialize rewriteLemmas : DiscrTreeCache (Name × Bool × Nat) ← unsafe do let path ← cachePath if (← path.pathExists) then - let (d, _r) ← unpickle (DiscrTree (Name × Bool × Nat)) path - -- We can drop the `CompactedRegion` value; we do not plan to free it - DiscrTreeCache.mk "rw?: using cache" processLemma (init := some d) + -- We can drop the `CompactedRegion` value from `unpickle`; we do not plan to free it + let d := (·.1) <$> unpickle (DiscrTree (Name × Bool × Nat)) path + DiscrTreeCache.mk "rw?: using cache" processLemma (init := d) else buildDiscrTree @@ -322,7 +320,7 @@ syntax forbidden := " [" (("-" ident),*,?) "]" `rw?` should not be left in proofs; it is a search tool, like `apply?`. -Suggestions are printed as `rw [h]` or `rw [←h]`. +Suggestions are printed as `rw [h]` or `rw [← h]`. You can use `rw? [-my_lemma, -my_theorem]` to prevent `rw?` using the named lemmas. -/ diff --git a/Mathlib/Tactic/Simps/Basic.lean b/Mathlib/Tactic/Simps/Basic.lean index 7cdceb6df4dc7..eaa82faf74ced 100644 --- a/Mathlib/Tactic/Simps/Basic.lean +++ b/Mathlib/Tactic/Simps/Basic.lean @@ -479,7 +479,7 @@ def findProjectionIndices (strName projName : Name) : MetaM (List Nat) := do partial def getCompositeOfProjectionsAux (proj : String) (e : Expr) (pos : Array Nat) (args : Array Expr) : MetaM (Expr × Array Nat) := do let env ← getEnv - let .const structName _ := (← whnf (←inferType e)).getAppFn | + let .const structName _ := (← whnf (← inferType e)).getAppFn | throwError "{e} doesn't have a structure as type" let projs := getStructureFieldsFlattened env structName let projInfo := projs.toList.map fun p ↦ do diff --git a/Mathlib/Tactic/ToAdditive.lean b/Mathlib/Tactic/ToAdditive.lean index 5f501628c5119..25f0ea907a8bc 100644 --- a/Mathlib/Tactic/ToAdditive.lean +++ b/Mathlib/Tactic/ToAdditive.lean @@ -734,6 +734,8 @@ def nameDict : String → List String | "units" => ["add", "Units"] | "cyclic" => ["add", "Cyclic"] | "rootable" => ["divisible"] + | "commute" => ["add", "Commute"] + | "semiconj" => ["add", "Semiconj"] | x => [x] /-- @@ -808,6 +810,10 @@ def fixAbbreviation : List String → List String | "Is"::"Of"::"Fin"::"Order"::s => "IsOfFinAddOrder" :: fixAbbreviation s | "is" :: "Central" :: "Scalar" :: s => "isCentralVAdd" :: fixAbbreviation s | "Is" :: "Central" :: "Scalar" :: s => "IsCentralVAdd" :: fixAbbreviation s + | "function" :: "_" :: "add" :: "Semiconj" :: s + => "function" :: "_" :: "semiconj" :: fixAbbreviation s + | "function" :: "_" :: "add" :: "Commute" :: s + => "function" :: "_" :: "commute" :: fixAbbreviation s | x :: s => x :: fixAbbreviation s | [] => [] diff --git a/Mathlib/Tactic/Widget/Conv.lean b/Mathlib/Tactic/Widget/Conv.lean index ce8855dccd772..c8df05dbd74a0 100644 --- a/Mathlib/Tactic/Widget/Conv.lean +++ b/Mathlib/Tactic/Widget/Conv.lean @@ -76,7 +76,7 @@ private def solveLevel (expr : Expr) (path : List Nat) : MetaM SolveReturn := ma | _ => do return { - expr := ←(Lean.Core.viewSubexpr path.head! expr) + expr := ← (Lean.Core.viewSubexpr path.head! expr) val? := toString (path.head! + 1) listRest := path.tail! } diff --git a/Mathlib/Testing/SlimCheck/Gen.lean b/Mathlib/Testing/SlimCheck/Gen.lean index 7da83f3ed8198..501fa156ee0fe 100644 --- a/Mathlib/Testing/SlimCheck/Gen.lean +++ b/Mathlib/Testing/SlimCheck/Gen.lean @@ -78,7 +78,7 @@ variable {α : Type u} /-- Create an `Array` of examples using `x`. The size is controlled by the size parameter of `Gen`. -/ def arrayOf (x : Gen α) : Gen (Array α) := do - let ⟨sz⟩ ← (ULiftable.up <| do choose Nat 0 (←getSize) (Nat.zero_le _) : Gen (ULift ℕ)) + let ⟨sz⟩ ← (ULiftable.up <| do choose Nat 0 (← getSize) (Nat.zero_le _) : Gen (ULift ℕ)) let mut res := #[] for _ in [0:sz] do res := res.push (← x) diff --git a/Mathlib/Testing/SlimCheck/Sampleable.lean b/Mathlib/Testing/SlimCheck/Sampleable.lean index 3808fac1df464..b580a2d1f7525 100644 --- a/Mathlib/Testing/SlimCheck/Sampleable.lean +++ b/Mathlib/Testing/SlimCheck/Sampleable.lean @@ -160,6 +160,10 @@ instance Fin.shrinkable {n : Nat} : Shrinkable (Fin n.succ) where instance Int.shrinkable : Shrinkable Int where shrink n := Nat.shrink n.natAbs |>.map (λ x => ([x, -x] : List ℤ)) |>.join +instance Rat.shrinkable : Shrinkable Rat where + shrink r := + (Shrinkable.shrink r.num).bind fun d => Nat.shrink r.den |>.map fun n => Rat.divInt d n + instance Bool.shrinkable : Shrinkable Bool := {} instance Char.shrinkable : Shrinkable Char := {} @@ -192,18 +196,25 @@ section Samplers open SampleableExt instance Nat.sampleableExt : SampleableExt Nat := - mkSelfContained (do choose Nat 0 (←getSize) (Nat.zero_le _)) + mkSelfContained (do choose Nat 0 (← getSize) (Nat.zero_le _)) instance Fin.sampleableExt {n : Nat} : SampleableExt (Fin (n.succ)) := - mkSelfContained (do choose (Fin n.succ) (Fin.ofNat 0) (Fin.ofNat (←getSize)) (by + mkSelfContained (do choose (Fin n.succ) (Fin.ofNat 0) (Fin.ofNat (← getSize)) (by simp only [LE.le, Fin.ofNat, Nat.zero_mod, Fin.zero_eta, Fin.val_zero, Nat.le_eq] exact Nat.zero_le _)) instance Int.sampleableExt : SampleableExt Int := mkSelfContained (do - choose Int (-(←getSize)) (←getSize) + choose Int (-(← getSize)) (← getSize) (le_trans (Int.neg_nonpos_of_nonneg (Int.ofNat_zero_le _)) (Int.ofNat_zero_le _))) +instance Rat.sampleableExt : SampleableExt Rat := + mkSelfContained (do + let d ← choose Int (-(← getSize)) (← getSize) + (le_trans (Int.neg_nonpos_of_nonneg (Int.ofNat_zero_le _)) (Int.ofNat_zero_le _)) + let n ← choose Nat 0 (← getSize) (Nat.zero_le _) + return Rat.divInt d n) + instance Bool.sampleableExt : SampleableExt Bool := mkSelfContained $ chooseAny Bool diff --git a/Mathlib/Testing/SlimCheck/Testable.lean b/Mathlib/Testing/SlimCheck/Testable.lean index b96dc6deddd5f..311e1e5492b08 100644 --- a/Mathlib/Testing/SlimCheck/Testable.lean +++ b/Mathlib/Testing/SlimCheck/Testable.lean @@ -334,7 +334,7 @@ partial def minimizeAux [SampleableExt α] {β : α → Prop} [∀ x, Testable ( slimTrace s!"{var} shrunk to {repr candidate} from {repr x}" let currentStep := OptionT.lift $ pure $ Sigma.mk candidate (addShrinks (n + 1) res) let nextStep := minimizeAux cfg var candidate (n + 1) - return ←(nextStep <|> currentStep) + return ← (nextStep <|> currentStep) if cfg.traceShrink then slimTrace s!"No shrinking possible for {var} := {repr x}" failure diff --git a/Mathlib/Topology/AlexandrovDiscrete.lean b/Mathlib/Topology/AlexandrovDiscrete.lean index e26cb73c66894..f708cf210a1c1 100644 --- a/Mathlib/Topology/AlexandrovDiscrete.lean +++ b/Mathlib/Topology/AlexandrovDiscrete.lean @@ -69,7 +69,7 @@ lemma isOpen_iInter₂ {f : ∀ i, κ i → Set α} (hf : ∀ i j, IsOpen (f i j isOpen_iInter fun _ ↦ isOpen_iInter $ hf _ lemma isClosed_sUnion (hS : ∀ s ∈ S, IsClosed s) : IsClosed (⋃₀ S) := by - simp only [←isOpen_compl_iff, compl_sUnion] at hS ⊢; exact isOpen_sInter $ ball_image_iff.2 hS + simp only [← isOpen_compl_iff, compl_sUnion] at hS ⊢; exact isOpen_sInter $ ball_image_iff.2 hS lemma isClosed_iUnion (hf : ∀ i, IsClosed (f i)) : IsClosed (⋃ i, f i) := isClosed_sUnion $ forall_range_iff.2 hf @@ -107,7 +107,7 @@ lemma interior_sInter (S : Set (Set α)) : interior (⋂₀ S) = ⋂ s ∈ S, in lemma closure_iUnion (f : ι → Set α) : closure (⋃ i, f i) = ⋃ i, closure (f i) := compl_injective <| by - simpa only [←interior_compl, compl_iUnion] using interior_iInter fun i ↦ (f i)ᶜ + simpa only [← interior_compl, compl_iUnion] using interior_iInter fun i ↦ (f i)ᶜ lemma closure_sUnion (S : Set (Set α)) : closure (⋃₀ S) = ⋃ s ∈ S, closure s := by simp_rw [sUnion_eq_biUnion, closure_iUnion] @@ -189,14 +189,14 @@ lemma gc_exterior_interior : GaloisConnection (exterior : Set α → Set α) int gc_exterior_interior.l_sup @[simp] lemma nhdsSet_exterior (s : Set α) : 𝓝ˢ (exterior s) = 𝓝ˢ s := by - ext t; simp_rw [←exterior_subset_iff_mem_nhdsSet, exterior_exterior] + ext t; simp_rw [← exterior_subset_iff_mem_nhdsSet, exterior_exterior] @[simp] lemma principal_exterior (s : Set α) : 𝓟 (exterior s) = 𝓝ˢ s := by - rw [←nhdsSet_exterior, isOpen_exterior.nhdsSet_eq] + rw [← nhdsSet_exterior, isOpen_exterior.nhdsSet_eq] @[simp] lemma exterior_subset_exterior : exterior s ⊆ exterior t ↔ 𝓝ˢ s ≤ 𝓝ˢ t := by refine ⟨?_, fun h ↦ ker_mono h⟩ - simp_rw [le_def, ←exterior_subset_iff_mem_nhdsSet] + simp_rw [le_def, ← exterior_subset_iff_mem_nhdsSet] exact fun h u ↦ h.trans lemma specializes_iff_exterior_subset : x ⤳ y ↔ exterior {x} ⊆ exterior {y} := by diff --git a/Mathlib/Topology/Algebra/Group/Basic.lean b/Mathlib/Topology/Algebra/Group/Basic.lean index 58ab47e982c0b..ce7db588365a7 100644 --- a/Mathlib/Topology/Algebra/Group/Basic.lean +++ b/Mathlib/Topology/Algebra/Group/Basic.lean @@ -145,7 +145,8 @@ theorem IsClosed.rightCoset {U : Set G} (h : IsClosed U) (x : G) : IsClosed (rig #align is_closed.right_add_coset IsClosed.right_addCoset @[to_additive] -theorem discreteTopology_of_open_singleton_one (h : IsOpen ({1} : Set G)) : DiscreteTopology G := by +theorem discreteTopology_of_isOpen_singleton_one (h : IsOpen ({1} : Set G)) : + DiscreteTopology G := by rw [← singletons_open_iff_discrete] intro g suffices {g} = (fun x : G => g⁻¹ * x) ⁻¹' {1} by @@ -153,14 +154,14 @@ theorem discreteTopology_of_open_singleton_one (h : IsOpen ({1} : Set G)) : Disc exact (continuous_mul_left g⁻¹).isOpen_preimage _ h simp only [mul_one, Set.preimage_mul_left_singleton, eq_self_iff_true, inv_inv, Set.singleton_eq_singleton_iff] -#align discrete_topology_of_open_singleton_one discreteTopology_of_open_singleton_one -#align discrete_topology_of_open_singleton_zero discreteTopology_of_open_singleton_zero +#align discrete_topology_of_open_singleton_one discreteTopology_of_isOpen_singleton_one +#align discrete_topology_of_open_singleton_zero discreteTopology_of_isOpen_singleton_zero @[to_additive] -theorem discreteTopology_iff_open_singleton_one : DiscreteTopology G ↔ IsOpen ({1} : Set G) := - ⟨fun h => forall_open_iff_discrete.mpr h {1}, discreteTopology_of_open_singleton_one⟩ -#align discrete_topology_iff_open_singleton_one discreteTopology_iff_open_singleton_one -#align discrete_topology_iff_open_singleton_zero discreteTopology_iff_open_singleton_zero +theorem discreteTopology_iff_isOpen_singleton_one : DiscreteTopology G ↔ IsOpen ({1} : Set G) := + ⟨fun h => forall_open_iff_discrete.mpr h {1}, discreteTopology_of_isOpen_singleton_one⟩ +#align discrete_topology_iff_open_singleton_one discreteTopology_iff_isOpen_singleton_one +#align discrete_topology_iff_open_singleton_zero discreteTopology_iff_isOpen_singleton_zero end ContinuousMulGroup diff --git a/Mathlib/Topology/Algebra/InfiniteSum/Basic.lean b/Mathlib/Topology/Algebra/InfiniteSum/Basic.lean index dc430ad6c00d9..95ebda29f1eea 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/Basic.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/Basic.lean @@ -619,6 +619,11 @@ theorem tsum_eq_tsum_of_ne_zero_bij {g : γ → α} (i : support g → β) tsum_eq_tsum_of_hasSum_iff_hasSum (hasSum_iff_hasSum_of_ne_zero_bij i hi hf hfg) #align tsum_eq_tsum_of_ne_zero_bij tsum_eq_tsum_of_ne_zero_bij +@[simp] +lemma tsum_extend_zero {γ : Type*} {g : γ → β} (hg : Injective g) (f : γ → α) : + ∑' y, extend g f 0 y = ∑' x, f x := + tsum_eq_tsum_of_hasSum_iff_hasSum <| hasSum_extend_zero hg + /-! ### `tsum` on subsets -/ theorem tsum_subtype (s : Set β) (f : β → α) : ∑' x : s, f x = ∑' x, s.indicator f x := diff --git a/Mathlib/Topology/Algebra/Module/Alternating.lean b/Mathlib/Topology/Algebra/Module/Alternating/Basic.lean similarity index 99% rename from Mathlib/Topology/Algebra/Module/Alternating.lean rename to Mathlib/Topology/Algebra/Module/Alternating/Basic.lean index af0a1451e9a99..a8cf78c2243d0 100644 --- a/Mathlib/Topology/Algebra/Module/Alternating.lean +++ b/Mathlib/Topology/Algebra/Module/Alternating/Basic.lean @@ -5,7 +5,7 @@ Authors: Yury Kudryashov, Heather Macbeth, Sébastien Gouëzel -/ import Mathlib.LinearAlgebra.Alternating.Basic import Mathlib.LinearAlgebra.BilinearMap -import Mathlib.Topology.Algebra.Module.Multilinear +import Mathlib.Topology.Algebra.Module.Multilinear.Basic /-! # Continuous alternating multilinear maps @@ -229,7 +229,7 @@ def applyAddHom (v : ι → M) : M [Λ^ι]→L[R] N →+ N := @[simp] theorem sum_apply {α : Type*} (f : α → M [Λ^ι]→L[R] N) (m : ι → M) {s : Finset α} : (∑ a in s, f a) m = ∑ a in s, f a m := - (applyAddHom m).map_sum f s + map_sum (applyAddHom m) f s /-- Projection to `ContinuousMultilinearMap`s as a bundled `AddMonoidHom`. -/ @[simps] diff --git a/Mathlib/Topology/Algebra/Module/Basic.lean b/Mathlib/Topology/Algebra/Module/Basic.lean index b13acdfeb7107..72e581dbfc63a 100644 --- a/Mathlib/Topology/Algebra/Module/Basic.lean +++ b/Mathlib/Topology/Algebra/Module/Basic.lean @@ -667,7 +667,7 @@ instance uniqueOfRight [Subsingleton M₂] : Unique (M₁ →SL[σ₁₂] M₂) #align continuous_linear_map.unique_of_right ContinuousLinearMap.uniqueOfRight theorem exists_ne_zero {f : M₁ →SL[σ₁₂] M₂} (hf : f ≠ 0) : ∃ x, f x ≠ 0 := by - by_contra' h + by_contra! h exact hf (ContinuousLinearMap.ext h) #align continuous_linear_map.exists_ne_zero ContinuousLinearMap.exists_ne_zero @@ -768,7 +768,7 @@ instance addCommMonoid : AddCommMonoid (M₁ →SL[σ₁₂] M₂) where @[simp, norm_cast] theorem coe_sum {ι : Type*} (t : Finset ι) (f : ι → M₁ →SL[σ₁₂] M₂) : ↑(∑ d in t, f d) = (∑ d in t, f d : M₁ →ₛₗ[σ₁₂] M₂) := - (AddMonoidHom.mk ⟨((↑) : (M₁ →SL[σ₁₂] M₂) → M₁ →ₛₗ[σ₁₂] M₂), rfl⟩ fun _ _ => rfl).map_sum _ _ + map_sum (AddMonoidHom.mk ⟨((↑) : (M₁ →SL[σ₁₂] M₂) → M₁ →ₛₗ[σ₁₂] M₂), rfl⟩ fun _ _ => rfl) _ _ #align continuous_linear_map.coe_sum ContinuousLinearMap.coe_sum @[simp, norm_cast] @@ -942,9 +942,9 @@ instance applySMulCommClass' : SMulCommClass (M₁ →L[R₁] M₁) R₁ M₁ wh smul_comm := ContinuousLinearMap.map_smul #align continuous_linear_map.apply_smul_comm_class' ContinuousLinearMap.applySMulCommClass' -instance continuousConstSMul : ContinuousConstSMul (M₁ →L[R₁] M₁) M₁ := +instance continuousConstSMul_apply : ContinuousConstSMul (M₁ →L[R₁] M₁) M₁ := ⟨ContinuousLinearMap.continuous⟩ -#align continuous_linear_map.has_continuous_const_smul ContinuousLinearMap.continuousConstSMul +#align continuous_linear_map.has_continuous_const_smul ContinuousLinearMap.continuousConstSMul_apply end ApplyAction @@ -2672,11 +2672,11 @@ def ClosedComplemented (p : Submodule R M) : Prop := ∃ f : M →L[R] p, ∀ x : p, f x = x #align submodule.closed_complemented Submodule.ClosedComplemented -theorem ClosedComplemented.has_closed_complement {p : Submodule R M} [T1Space p] +theorem ClosedComplemented.exists_isClosed_isCompl {p : Submodule R M} [T1Space p] (h : ClosedComplemented p) : - ∃ (q : Submodule R M) (_ : IsClosed (q : Set M)), IsCompl p q := + ∃ q : Submodule R M, IsClosed (q : Set M) ∧ IsCompl p q := Exists.elim h fun f hf => ⟨ker f, isClosed_ker f, LinearMap.isCompl_of_proj hf⟩ -#align submodule.closed_complemented.has_closed_complement Submodule.ClosedComplemented.has_closed_complement +#align submodule.closed_complemented.has_closed_complement Submodule.ClosedComplemented.exists_isClosed_isCompl protected theorem ClosedComplemented.isClosed [TopologicalAddGroup M] [T1Space M] {p : Submodule R M} (h : ClosedComplemented p) : IsClosed (p : Set M) := by @@ -2695,6 +2695,20 @@ theorem closedComplemented_top : ClosedComplemented (⊤ : Submodule R M) := ⟨(id R M).codRestrict ⊤ fun _x => trivial, fun x => Subtype.ext_iff_val.2 <| by simp⟩ #align submodule.closed_complemented_top Submodule.closedComplemented_top +/-- If `p` is a closed complemented submodule, +then there exists a submodule `q` and a continuous linear equivalence `M ≃L[R] (p × q)` such that +`e (x : p) = (x, 0)`, `e (y : q) = (0, y)`, and `e.symm x = x.1 + x.2`. + +In fact, the properties of `e` imply the properties of `e.symm` and vice versa, +but we provide both for convenience. -/ +lemma ClosedComplemented.exists_submodule_equiv_prod [TopologicalAddGroup M] + {p : Submodule R M} (hp : p.ClosedComplemented) : + ∃ (q : Submodule R M) (e : M ≃L[R] (p × q)), + (∀ x : p, e x = (x, 0)) ∧ (∀ y : q, e y = (0, y)) ∧ (∀ x, e.symm x = x.1 + x.2) := + let ⟨f, hf⟩ := hp + ⟨LinearMap.ker f, .equivOfRightInverse _ p.subtypeL hf, + fun _ ↦ by ext <;> simp [hf], fun _ ↦ by ext <;> simp [hf], fun _ ↦ rfl⟩ + end Submodule theorem ContinuousLinearMap.closedComplemented_ker_of_rightInverse {R : Type*} [Ring R] diff --git a/Mathlib/Topology/Algebra/Module/FiniteDimension.lean b/Mathlib/Topology/Algebra/Module/FiniteDimension.lean index 455f74ca9b138..afb64a2a6e7f9 100644 --- a/Mathlib/Topology/Algebra/Module/FiniteDimension.lean +++ b/Mathlib/Topology/Algebra/Module/FiniteDimension.lean @@ -100,7 +100,7 @@ theorem unique_topology_of_t2 {t : TopologicalSpace 𝕜} (h₁ : @TopologicalAd · rw [mem_closedBall_zero_iff] -- Now suppose `ξ ≠ 0`. By contradiction, let's assume `ε < ‖ξ‖`, and show that -- `ξ₀ ∈ 𝓑 ⊆ {ξ₀}ᶜ`, which is a contradiction. - by_contra' h + by_contra! h suffices (ξ₀ * ξ⁻¹) • ξ ∈ balancedCore 𝕜 {ξ₀}ᶜ by rw [smul_eq_mul 𝕜, mul_assoc, inv_mul_cancel hξ0, mul_one] at this exact not_mem_compl_iff.mpr (mem_singleton ξ₀) ((balancedCore_subset _) this) @@ -449,6 +449,7 @@ theorem coe_constrL (v : Basis ι 𝕜 E) (f : ι → F) : (v.constrL f : E → /-- The continuous linear equivalence between a vector space over `𝕜` with a finite basis and functions from its basis indexing type to `𝕜`. -/ +@[simps! apply] def equivFunL (v : Basis ι 𝕜 E) : E ≃L[𝕜] ι → 𝕜 := { v.equivFun with continuous_toFun := @@ -459,6 +460,11 @@ def equivFunL (v : Basis ι 𝕜 E) : E ≃L[𝕜] ι → 𝕜 := exact v.equivFun.symm.toLinearMap.continuous_of_finiteDimensional } #align basis.equiv_funL Basis.equivFunL +@[simp] +lemma equivFunL_symm_apply_repr (v : Basis ι 𝕜 E) (x : E) : + v.equivFunL.symm (v.repr x) = x := + v.equivFunL.symm_apply_apply x + @[simp] theorem constrL_apply (v : Basis ι 𝕜 E) (f : ι → F) (e : E) : v.constrL f e = ∑ i, v.equivFun e i • f i := diff --git a/Mathlib/Topology/Algebra/Module/Multilinear.lean b/Mathlib/Topology/Algebra/Module/Multilinear/Basic.lean similarity index 99% rename from Mathlib/Topology/Algebra/Module/Multilinear.lean rename to Mathlib/Topology/Algebra/Module/Multilinear/Basic.lean index 6f070431be56a..861554ebe6050 100644 --- a/Mathlib/Topology/Algebra/Module/Multilinear.lean +++ b/Mathlib/Topology/Algebra/Module/Multilinear/Basic.lean @@ -220,7 +220,7 @@ def applyAddHom (m : ∀ i, M₁ i) : ContinuousMultilinearMap R M₁ M₂ →+ @[simp] theorem sum_apply {α : Type*} (f : α → ContinuousMultilinearMap R M₁ M₂) (m : ∀ i, M₁ i) {s : Finset α} : (∑ a in s, f a) m = ∑ a in s, f a m := - (applyAddHom m).map_sum f s + map_sum (applyAddHom m) f s #align continuous_multilinear_map.sum_apply ContinuousMultilinearMap.sum_apply end ContinuousAdd diff --git a/Mathlib/Topology/Algebra/Module/StrongTopology.lean b/Mathlib/Topology/Algebra/Module/StrongTopology.lean index 02afed1729dff..07491213f2fc9 100644 --- a/Mathlib/Topology/Algebra/Module/StrongTopology.lean +++ b/Mathlib/Topology/Algebra/Module/StrongTopology.lean @@ -172,6 +172,21 @@ theorem strongTopology.hasBasis_nhds_zero [TopologicalSpace F] [TopologicalAddGr strongTopology.hasBasis_nhds_zero_of_basis σ F 𝔖 h𝔖₁ h𝔖₂ (𝓝 0).basis_sets #align continuous_linear_map.strong_topology.has_basis_nhds_zero ContinuousLinearMap.strongTopology.hasBasis_nhds_zero +theorem strongTopology.continuousConstSMul {M : Type*} + [Monoid M] [DistribMulAction M F] [SMulCommClass 𝕜₂ M F] + [TopologicalSpace F] [TopologicalAddGroup F] [ContinuousConstSMul M F] (𝔖 : Set (Set E)) + (h𝔖₁ : 𝔖.Nonempty) (h𝔖₂ : DirectedOn (· ⊆ ·) 𝔖) : + @ContinuousConstSMul M (E →SL[σ] F) (strongTopology σ F 𝔖) _ := by + letI := strongTopology σ F 𝔖 + haveI : TopologicalAddGroup (E →SL[σ] F) := strongTopology.topologicalAddGroup σ F 𝔖 + refine ⟨fun c ↦ continuous_of_continuousAt_zero (DistribSMul.toAddMonoidHom _ c) ?_⟩ + have H₁ := strongTopology.hasBasis_nhds_zero σ F _ h𝔖₁ h𝔖₂ + have H₂ : Filter.Tendsto (c • ·) (𝓝 0 : Filter F) (𝓝 0) := + (continuous_const_smul c).tendsto' 0 _ (smul_zero _) + rw [ContinuousAt, map_zero, H₁.tendsto_iff H₁] + rintro ⟨s, t⟩ ⟨hs : s ∈ 𝔖, ht : t ∈ 𝓝 0⟩ + exact ⟨(s, (c • ·) ⁻¹' t), ⟨hs, H₂ ht⟩, fun f ↦ _root_.id⟩ + end General section BoundedSets @@ -225,6 +240,13 @@ protected theorem hasBasis_nhds_zero [TopologicalSpace F] [TopologicalAddGroup F ContinuousLinearMap.hasBasis_nhds_zero_of_basis (𝓝 0).basis_sets #align continuous_linear_map.has_basis_nhds_zero ContinuousLinearMap.hasBasis_nhds_zero +instance continuousConstSMul {M : Type*} [Monoid M] [DistribMulAction M F] [SMulCommClass 𝕜₂ M F] + [TopologicalSpace F] [TopologicalAddGroup F] [ContinuousConstSMul M F] : + ContinuousConstSMul M (E →SL[σ] F) := + strongTopology.continuousConstSMul σ F {S | Bornology.IsVonNBounded 𝕜₁ S} + ⟨∅, Bornology.isVonNBounded_empty 𝕜₁ E⟩ + (directedOn_of_sup_mem fun _ _ => Bornology.IsVonNBounded.union) + variable (G) [TopologicalSpace F] [TopologicalSpace G] /-- Pre-composition by a *fixed* continuous linear map as a continuous linear map. diff --git a/Mathlib/Topology/Algebra/Nonarchimedean/AdicTopology.lean b/Mathlib/Topology/Algebra/Nonarchimedean/AdicTopology.lean index 9a23271b177f5..df8353733e861 100644 --- a/Mathlib/Topology/Algebra/Nonarchimedean/AdicTopology.lean +++ b/Mathlib/Topology/Algebra/Nonarchimedean/AdicTopology.lean @@ -210,7 +210,7 @@ theorem is_bot_adic_iff {A : Type*} [CommRing A] [TopologicalSpace A] [Topologic rw [isAdic_iff] constructor · rintro ⟨h, _h'⟩ - rw [discreteTopology_iff_open_singleton_zero] + rw [discreteTopology_iff_isOpen_singleton_zero] simpa using h 1 · intros constructor diff --git a/Mathlib/Topology/Algebra/OpenSubgroup.lean b/Mathlib/Topology/Algebra/OpenSubgroup.lean index 9f41fe48d299d..208e0e1b86ca8 100644 --- a/Mathlib/Topology/Algebra/OpenSubgroup.lean +++ b/Mathlib/Topology/Algebra/OpenSubgroup.lean @@ -383,9 +383,9 @@ variable {R : Type*} [CommRing R] variable [TopologicalSpace R] [TopologicalRing R] -theorem isOpen_of_open_subideal {U I : Ideal R} (h : U ≤ I) (hU : IsOpen (U : Set R)) : +theorem isOpen_of_isOpen_subideal {U I : Ideal R} (h : U ≤ I) (hU : IsOpen (U : Set R)) : IsOpen (I : Set R) := @Submodule.isOpen_mono R R _ _ _ _ Semiring.toModule _ _ h hU -#align ideal.is_open_of_open_subideal Ideal.isOpen_of_open_subideal +#align ideal.is_open_of_open_subideal Ideal.isOpen_of_isOpen_subideal end Ideal diff --git a/Mathlib/Topology/Algebra/Order/LiminfLimsup.lean b/Mathlib/Topology/Algebra/Order/LiminfLimsup.lean index 13bdb361e1f54..79654909e4d9a 100644 --- a/Mathlib/Topology/Algebra/Order/LiminfLimsup.lean +++ b/Mathlib/Topology/Algebra/Order/LiminfLimsup.lean @@ -263,7 +263,7 @@ theorem tendsto_of_no_upcrossings [DenselyOrdered α] {f : Filter β} {u : β · exact ⟨sInf ∅, tendsto_bot⟩ refine' ⟨limsup u f, _⟩ apply tendsto_of_le_liminf_of_limsup_le _ le_rfl h h' - by_contra' hlt + by_contra! hlt obtain ⟨a, ⟨⟨la, au⟩, as⟩⟩ : ∃ a, (f.liminf u < a ∧ a < f.limsup u) ∧ a ∈ s := dense_iff_inter_open.1 hs (Set.Ioo (f.liminf u) (f.limsup u)) isOpen_Ioo (Set.nonempty_Ioo.2 hlt) @@ -357,13 +357,13 @@ theorem Antitone.map_limsSup_of_continuousAt {F : Filter R} [NeBot F] {f : R → have B : ∃ᶠ n in F, F.limsSup ≤ n := by apply (frequently_lt_of_lt_limsSup cobdd c_lt).mono intro x hx - by_contra' + by_contra! have : (Set.Ioo c F.limsSup).Nonempty := ⟨x, ⟨hx, this⟩⟩ simp only [hc, Set.not_nonempty_empty] at this apply liminf_le_of_frequently_le _ (bdd_above.isBoundedUnder f_decr) exact (B.mono fun x hx ↦ f_decr hx) push_neg at h' - by_contra' H + by_contra! H have not_bot : ¬ IsBot F.limsSup := fun maybe_bot ↦ lt_irrefl (F.liminf f) <| lt_of_le_of_lt (liminf_le_of_frequently_le (frequently_of_forall (fun r ↦ f_decr (maybe_bot r))) @@ -546,8 +546,7 @@ theorem limsup_eq_tendsto_sum_indicator_nat_atTop (s : ℕ → Set α) : exact zero_lt_one · rintro hω i rw [Set.mem_setOf_eq, tendsto_atTop_atTop] at hω - by_contra hcon - push_neg at hcon + by_contra! hcon obtain ⟨j, h⟩ := hω (i + 1) have : (∑ k in Finset.range j, (s (k + 1)).indicator 1 ω) ≤ i := by have hle : ∀ j ≤ i, (∑ k in Finset.range j, (s (k + 1)).indicator 1 ω) ≤ i := by diff --git a/Mathlib/Topology/Algebra/UniformGroup.lean b/Mathlib/Topology/Algebra/UniformGroup.lean index e5e88611b6312..4c4c8238c946c 100644 --- a/Mathlib/Topology/Algebra/UniformGroup.lean +++ b/Mathlib/Topology/Algebra/UniformGroup.lean @@ -686,7 +686,7 @@ theorem comm_topologicalGroup_is_uniform : UniformGroup G := by tendsto_comap_iff, prod_comap_comap_eq] simp only [Function.comp, div_eq_mul_inv, mul_inv_rev, inv_inv, mul_comm, mul_left_comm] at * simp only [inv_one, mul_one, ← mul_assoc] at this - simp_rw [←mul_assoc, mul_comm] + simp_rw [← mul_assoc, mul_comm] assumption #align topological_comm_group_is_uniform comm_topologicalGroup_is_uniform #align topological_add_comm_group_is_uniform comm_topologicalAddGroup_is_uniform diff --git a/Mathlib/Topology/Bases.lean b/Mathlib/Topology/Bases.lean index c25bbf94726ce..073e5e7270259 100644 --- a/Mathlib/Topology/Bases.lean +++ b/Mathlib/Topology/Bases.lean @@ -124,7 +124,7 @@ theorem isTopologicalBasis_of_subbasis {s : Set (Set α)} (hs : t = generateFrom /-- If a family of open sets `s` is such that every open neighbourhood contains some member of `s`, then `s` is a topological basis. -/ -theorem isTopologicalBasis_of_open_of_nhds {s : Set (Set α)} (h_open : ∀ u ∈ s, IsOpen u) +theorem isTopologicalBasis_of_isOpen_of_nhds {s : Set (Set α)} (h_open : ∀ u ∈ s, IsOpen u) (h_nhds : ∀ (a : α) (u : Set α), a ∈ u → IsOpen u → ∃ v ∈ s, a ∈ v ∧ v ⊆ u) : IsTopologicalBasis s := by refine' @@ -137,7 +137,7 @@ theorem isTopologicalBasis_of_open_of_nhds {s : Set (Set α)} (h_open : ∀ u rcases h_nhds a u ha hu with ⟨v, hvs, hav, hvu⟩ rw [nhds_generateFrom] exact iInf₂_le_of_le v ⟨hav, hvs⟩ (le_principal_iff.2 hvu) -#align topological_space.is_topological_basis_of_open_of_nhds TopologicalSpace.isTopologicalBasis_of_open_of_nhds +#align topological_space.is_topological_basis_of_open_of_nhds TopologicalSpace.isTopologicalBasis_of_isOpen_of_nhds /-- A set `s` is in the neighbourhood of `a` iff there is some basis set `t`, which contains `a` and is itself contained in `s`. -/ @@ -245,13 +245,13 @@ theorem IsTopologicalBasis.exists_nonempty_subset {B : Set (Set α)} (hb : IsTop #align topological_space.is_topological_basis.exists_nonempty_subset TopologicalSpace.IsTopologicalBasis.exists_nonempty_subset theorem isTopologicalBasis_opens : IsTopologicalBasis { U : Set α | IsOpen U } := - isTopologicalBasis_of_open_of_nhds (by tauto) (by tauto) + isTopologicalBasis_of_isOpen_of_nhds (by tauto) (by tauto) #align topological_space.is_topological_basis_opens TopologicalSpace.isTopologicalBasis_opens protected theorem IsTopologicalBasis.prod {β} [TopologicalSpace β] {B₁ : Set (Set α)} {B₂ : Set (Set β)} (h₁ : IsTopologicalBasis B₁) (h₂ : IsTopologicalBasis B₂) : IsTopologicalBasis (image2 (· ×ˢ ·) B₁ B₂) := by - refine' isTopologicalBasis_of_open_of_nhds _ _ + refine' isTopologicalBasis_of_isOpen_of_nhds _ _ · rintro _ ⟨u₁, u₂, hu₁, hu₂, rfl⟩ exact (h₁.isOpen hu₁).prod (h₂.isOpen hu₂) · rintro ⟨a, b⟩ u hu uo @@ -262,7 +262,7 @@ protected theorem IsTopologicalBasis.prod {β} [TopologicalSpace β] {B₁ : Set protected theorem IsTopologicalBasis.inducing {β} [TopologicalSpace β] {f : α → β} {T : Set (Set β)} (hf : Inducing f) (h : IsTopologicalBasis T) : IsTopologicalBasis ((preimage f) '' T) := by - refine' isTopologicalBasis_of_open_of_nhds _ _ + refine' isTopologicalBasis_of_isOpen_of_nhds _ _ · rintro _ ⟨V, hV, rfl⟩ rw [hf.isOpen_iff] refine' ⟨V, h.isOpen hV, rfl⟩ @@ -277,7 +277,7 @@ protected theorem IsTopologicalBasis.inducing {β} [TopologicalSpace β] {f : α theorem isTopologicalBasis_of_cover {ι} {U : ι → Set α} (Uo : ∀ i, IsOpen (U i)) (Uc : ⋃ i, U i = univ) {b : ∀ i, Set (Set (U i))} (hb : ∀ i, IsTopologicalBasis (b i)) : IsTopologicalBasis (⋃ i : ι, image ((↑) : U i → α) '' b i) := by - refine' isTopologicalBasis_of_open_of_nhds (fun u hu => _) _ + refine' isTopologicalBasis_of_isOpen_of_nhds (fun u hu => _) _ · simp only [mem_iUnion, mem_image] at hu rcases hu with ⟨i, s, sb, rfl⟩ exact (Uo i).isOpenMap_subtype_val _ ((hb i).isOpen sb) @@ -528,7 +528,7 @@ theorem isTopologicalBasis_pi {ι : Type*} {X : ι → Type*} [∀ i, Topologica {T : ∀ i, Set (Set (X i))} (cond : ∀ i, IsTopologicalBasis (T i)) : IsTopologicalBasis { S | ∃ (U : ∀ i, Set (X i)) (F : Finset ι), (∀ i, i ∈ F → U i ∈ T i) ∧ S = (F : Set ι).pi U } := by - refine' isTopologicalBasis_of_open_of_nhds _ _ + refine' isTopologicalBasis_of_isOpen_of_nhds _ _ · rintro _ ⟨U, F, h1, rfl⟩ apply isOpen_set_pi F.finite_toSet intro i hi @@ -565,7 +565,7 @@ theorem isTopologicalBasis_iInf {β : Type*} {ι : Type*} {X : ι → Type*} theorem isTopologicalBasis_singletons (α : Type*) [TopologicalSpace α] [DiscreteTopology α] : IsTopologicalBasis { s | ∃ x : α, (s : Set α) = {x} } := - isTopologicalBasis_of_open_of_nhds (fun _ _ => isOpen_discrete _) fun x _ hx _ => + isTopologicalBasis_of_isOpen_of_nhds (fun _ _ => isOpen_discrete _) fun x _ hx _ => ⟨{x}, ⟨x, rfl⟩, mem_singleton x, singleton_subset_iff.2 hx⟩ #align is_topological_basis_singletons isTopologicalBasis_singletons @@ -841,7 +841,7 @@ topological bases on each of the parts of the space. -/ theorem IsTopologicalBasis.sigma {s : ∀ i : ι, Set (Set (E i))} (hs : ∀ i, IsTopologicalBasis (s i)) : IsTopologicalBasis (⋃ i : ι, (fun u => (Sigma.mk i '' u : Set (Σi, E i))) '' s i) := by - apply isTopologicalBasis_of_open_of_nhds + apply isTopologicalBasis_of_isOpen_of_nhds · intro u hu obtain ⟨i, t, ts, rfl⟩ : ∃ (i : ι) (t : Set (E i)), t ∈ s i ∧ Sigma.mk i '' t = u := by simpa only [mem_iUnion, mem_image] using hu @@ -874,7 +874,7 @@ topological bases on each of the two components. -/ theorem IsTopologicalBasis.sum {s : Set (Set α)} (hs : IsTopologicalBasis s) {t : Set (Set β)} (ht : IsTopologicalBasis t) : IsTopologicalBasis ((fun u => Sum.inl '' u) '' s ∪ (fun u => Sum.inr '' u) '' t) := by - apply isTopologicalBasis_of_open_of_nhds + apply isTopologicalBasis_of_isOpen_of_nhds · rintro u (⟨w, hw, rfl⟩ | ⟨w, hw, rfl⟩) · exact openEmbedding_inl.isOpenMap w (hs.isOpen hw) · exact openEmbedding_inr.isOpenMap w (ht.isOpen hw) @@ -909,7 +909,7 @@ variable {X : Type*} [TopologicalSpace X] {Y : Type*} [TopologicalSpace Y] {π : /-- The image of a topological basis under an open quotient map is a topological basis. -/ theorem IsTopologicalBasis.quotientMap {V : Set (Set X)} (hV : IsTopologicalBasis V) (h' : QuotientMap π) (h : IsOpenMap π) : IsTopologicalBasis (Set.image π '' V) := by - apply isTopologicalBasis_of_open_of_nhds + apply isTopologicalBasis_of_isOpen_of_nhds · rintro - ⟨U, U_in_V, rfl⟩ apply h U (hV.isOpen U_in_V) · intro y U y_in_U U_open diff --git a/Mathlib/Topology/Basic.lean b/Mathlib/Topology/Basic.lean index 96355bb3caabc..2bf880c9f7dc1 100644 --- a/Mathlib/Topology/Basic.lean +++ b/Mathlib/Topology/Basic.lean @@ -1457,16 +1457,16 @@ theorem mem_closure_of_mem_closure_union {s₁ s₂ : Set α} {x : α} (h : x #align mem_closure_of_mem_closure_union mem_closure_of_mem_closure_union /-- The intersection of an open dense set with a dense set is a dense set. -/ -theorem Dense.inter_of_open_left {s t : Set α} (hs : Dense s) (ht : Dense t) (hso : IsOpen s) : +theorem Dense.inter_of_isOpen_left {s t : Set α} (hs : Dense s) (ht : Dense t) (hso : IsOpen s) : Dense (s ∩ t) := fun x => closure_minimal hso.inter_closure isClosed_closure <| by simp [hs.closure_eq, ht.closure_eq] -#align dense.inter_of_open_left Dense.inter_of_open_left +#align dense.inter_of_open_left Dense.inter_of_isOpen_left /-- The intersection of a dense set with an open dense set is a dense set. -/ -theorem Dense.inter_of_open_right {s t : Set α} (hs : Dense s) (ht : Dense t) (hto : IsOpen t) : +theorem Dense.inter_of_isOpen_right {s t : Set α} (hs : Dense s) (ht : Dense t) (hto : IsOpen t) : Dense (s ∩ t) := - inter_comm t s ▸ ht.inter_of_open_left hs hto -#align dense.inter_of_open_right Dense.inter_of_open_right + inter_comm t s ▸ ht.inter_of_isOpen_left hs hto +#align dense.inter_of_open_right Dense.inter_of_isOpen_right theorem Dense.inter_nhds_nonempty {s t : Set α} (hs : Dense s) {x : α} (ht : t ∈ 𝓝 x) : (s ∩ t).Nonempty := diff --git a/Mathlib/Topology/Category/Profinite/Basic.lean b/Mathlib/Topology/Category/Profinite/Basic.lean index af8cebc9db5cc..6f0ae6d4ed957 100644 --- a/Mathlib/Topology/Category/Profinite/Basic.lean +++ b/Mathlib/Topology/Category/Profinite/Basic.lean @@ -379,7 +379,7 @@ theorem epi_iff_surjective {X Y : Profinite.{u}} (f : X ⟶ Y) : Epi f ↔ Funct rintro ⟨y', hy'⟩ exact hy y' hy' have hUy : U ∈ 𝓝 y := hC.compl_mem_nhds hyU - obtain ⟨V, hV, hyV, hVU⟩ := isTopologicalBasis_clopen.mem_nhds_iff.mp hUy + obtain ⟨V, hV, hyV, hVU⟩ := isTopologicalBasis_isClopen.mem_nhds_iff.mp hUy classical let Z := of (ULift.{u} <| Fin 2) let g : Y ⟶ Z := ⟨(LocallyConstant.ofIsClopen hV).map ULift.up, LocallyConstant.continuous _⟩ diff --git a/Mathlib/Topology/Category/Profinite/CofilteredLimit.lean b/Mathlib/Topology/Category/Profinite/CofilteredLimit.lean index fb41cbcb2a4c3..25a281e841988 100644 --- a/Mathlib/Topology/Category/Profinite/CofilteredLimit.lean +++ b/Mathlib/Topology/Category/Profinite/CofilteredLimit.lean @@ -58,7 +58,7 @@ theorem exists_isClopen_of_cofiltered {U : Set C.pt} (hC : IsLimit C) (hU : IsCl rotate_left · intro i change TopologicalSpace.IsTopologicalBasis {W : Set (F.obj i) | IsClopen W} - apply isTopologicalBasis_clopen + apply isTopologicalBasis_isClopen · rintro i j f V (hV : IsClopen _) exact ⟨hV.1.preimage ((F ⋙ toTopCat).map f).continuous, hV.2.preimage ((F ⋙ toTopCat).map f).continuous⟩ diff --git a/Mathlib/Topology/Category/Stonean/Basic.lean b/Mathlib/Topology/Category/Stonean/Basic.lean index d06ab7e3a8d7a..ce06f29a1c3bb 100644 --- a/Mathlib/Topology/Category/Stonean/Basic.lean +++ b/Mathlib/Topology/Category/Stonean/Basic.lean @@ -185,13 +185,13 @@ lemma epi_iff_surjective {X Y : Stonean} (f : X ⟶ Y) : refine ⟨?_, ConcreteCategory.epi_of_surjective _⟩ dsimp [Function.Surjective] intro h y - by_contra' hy + by_contra! hy let C := Set.range f have hC : IsClosed C := (isCompact_range f.continuous).isClosed let U := Cᶜ have hUy : U ∈ nhds y := by simp only [Set.mem_range, hy, exists_false, not_false_eq_true, hC.compl_mem_nhds] - obtain ⟨V, hV, hyV, hVU⟩ := isTopologicalBasis_clopen.mem_nhds_iff.mp hUy + obtain ⟨V, hV, hyV, hVU⟩ := isTopologicalBasis_isClopen.mem_nhds_iff.mp hUy classical let g : Y ⟶ mkFinite (ULift (Fin 2)) := ⟨(LocallyConstant.ofIsClopen hV).map ULift.up, LocallyConstant.continuous _⟩ diff --git a/Mathlib/Topology/Category/TopCat/Basic.lean b/Mathlib/Topology/Category/TopCat/Basic.lean index 001acfdc2c818..b534a4fb82620 100644 --- a/Mathlib/Topology/Category/TopCat/Basic.lean +++ b/Mathlib/Topology/Category/TopCat/Basic.lean @@ -168,7 +168,7 @@ set_option linter.uppercaseLean3 false in @[simp] theorem openEmbedding_iff_comp_isIso' {X Y Z : TopCat} (f : X ⟶ Y) (g : Y ⟶ Z) [IsIso g] : OpenEmbedding ((forget TopCat).map f ≫ (forget TopCat).map g) ↔ OpenEmbedding f := by - simp only [←Functor.map_comp] + simp only [← Functor.map_comp] exact openEmbedding_iff_comp_isIso f g -- Porting note: simpNF requested partially simped version below @@ -185,7 +185,7 @@ set_option linter.uppercaseLean3 false in @[simp] theorem openEmbedding_iff_isIso_comp' {X Y Z : TopCat} (f : X ⟶ Y) (g : Y ⟶ Z) [IsIso f] : OpenEmbedding ((forget TopCat).map f ≫ (forget TopCat).map g) ↔ OpenEmbedding g := by - simp only [←Functor.map_comp] + simp only [← Functor.map_comp] exact openEmbedding_iff_isIso_comp f g end TopCat diff --git a/Mathlib/Topology/Category/TopCat/Limits/Basic.lean b/Mathlib/Topology/Category/TopCat/Limits/Basic.lean index f381f6d7ab6a5..09ee3962e834c 100644 --- a/Mathlib/Topology/Category/TopCat/Limits/Basic.lean +++ b/Mathlib/Topology/Category/TopCat/Limits/Basic.lean @@ -112,7 +112,7 @@ def limitConeInfiIsLimit (F : J ⥤ TopCatMax.{v, u}) : IsLimit (limitConeInfi.{ (fun s => ⟨fun v => ⟨ fun j => (Functor.mapCone forget s).π.app j v, ?_⟩, ?_⟩) fun s => ?_ · dsimp [Functor.sections] intro _ _ _ - rw [←comp_apply', forget_map_eq_coe, ←s.π.naturality, forget_map_eq_coe] + rw [← comp_apply', forget_map_eq_coe, ← s.π.naturality, forget_map_eq_coe] dsimp rw [Category.id_comp] · exact diff --git a/Mathlib/Topology/Category/TopCat/Limits/Cofiltered.lean b/Mathlib/Topology/Category/TopCat/Limits/Cofiltered.lean index 53c75bf16c37b..45932ffaad2ed 100644 --- a/Mathlib/Topology/Category/TopCat/Limits/Cofiltered.lean +++ b/Mathlib/Topology/Category/TopCat/Limits/Cofiltered.lean @@ -121,7 +121,7 @@ theorem isTopologicalBasis_cofiltered_limit (T : ∀ j, Set (Set (F.obj j))) rw [dif_pos he, ← Set.preimage_comp] apply congrFun apply congrArg - rw [←coe_comp, D.w] + rw [← coe_comp, D.w] rfl #align Top.is_topological_basis_cofiltered_limit TopCat.isTopologicalBasis_cofiltered_limit diff --git a/Mathlib/Topology/Category/TopCat/Limits/Products.lean b/Mathlib/Topology/Category/TopCat/Limits/Products.lean index 07f12e4d224d9..9ed64c2c05aa0 100644 --- a/Mathlib/Topology/Category/TopCat/Limits/Products.lean +++ b/Mathlib/Topology/Category/TopCat/Limits/Products.lean @@ -133,7 +133,7 @@ theorem sigmaIsoSigma_hom_ι_apply {ι : Type v} (α : ι → TopCatMax.{v, u}) @[simp] theorem sigmaIsoSigma_inv_apply {ι : Type v} (α : ι → TopCatMax.{v, u}) (i : ι) (x : α i) : (sigmaIsoSigma α).inv ⟨i, x⟩ = (Sigma.ι α i : _) x := by - rw [← sigmaIsoSigma_hom_ι_apply, ← comp_app, ←comp_app, Iso.hom_inv_id, + rw [← sigmaIsoSigma_hom_ι_apply, ← comp_app, ← comp_app, Iso.hom_inv_id, Category.comp_id] #align Top.sigma_iso_sigma_inv_apply TopCat.sigmaIsoSigma_inv_apply diff --git a/Mathlib/Topology/Category/TopCat/Limits/Pullbacks.lean b/Mathlib/Topology/Category/TopCat/Limits/Pullbacks.lean index f173aa6e68e30..4313fe71a6f48 100644 --- a/Mathlib/Topology/Category/TopCat/Limits/Pullbacks.lean +++ b/Mathlib/Topology/Category/TopCat/Limits/Pullbacks.lean @@ -170,7 +170,7 @@ theorem range_pullback_to_prod {X Y Z : TopCat} (f : X ⟶ Z) (g : Y ⟶ Z) : use (pullbackIsoProdSubtype f g).inv ⟨⟨_, _⟩, h⟩ apply Concrete.limit_ext rintro ⟨⟨⟩⟩ <;> - rw [←comp_apply, ←comp_apply, limit.lift_π] <;> + rw [← comp_apply, ← comp_apply, limit.lift_π] <;> -- This used to be `simp` before leanprover/lean4#2644 aesop_cat #align Top.range_pullback_to_prod TopCat.range_pullback_to_prod @@ -220,7 +220,7 @@ theorem range_pullback_map {W X Y Z S T : TopCat} (f₁ : W ⟶ S) (f₂ : X ⟶ ext constructor · rintro ⟨y, rfl⟩ - simp only [Set.mem_inter_iff, Set.mem_preimage, Set.mem_range, ←comp_apply, limit.lift_π, + simp only [Set.mem_inter_iff, Set.mem_preimage, Set.mem_range, ← comp_apply, limit.lift_π, PullbackCone.mk_pt, PullbackCone.mk_π_app] simp only [comp_apply, exists_apply_eq_apply, and_self] rintro ⟨⟨x₁, hx₁⟩, ⟨x₂, hx₂⟩⟩ @@ -232,7 +232,7 @@ theorem range_pullback_map {W X Y Z S T : TopCat} (f₁ : W ⟶ S) (f₂ : X ⟶ use (pullbackIsoProdSubtype f₁ f₂).inv ⟨⟨x₁, x₂⟩, this⟩ apply Concrete.limit_ext rintro (_ | _ | _) <;> - simp only [←comp_apply, Category.assoc, limit.lift_π, PullbackCone.mk_π_app_one] + simp only [← comp_apply, Category.assoc, limit.lift_π, PullbackCone.mk_π_app_one] · simp only [cospan_one, pullbackIsoProdSubtype_inv_fst_assoc, comp_apply, pullbackFst_apply, hx₁] rw [← limit.w _ WalkingCospan.Hom.inl, cospan_map_inl, comp_apply (g := g₁)] @@ -335,7 +335,7 @@ theorem embedding_of_pullback_embeddings {X Y S : TopCat} {f : X ⟶ S} {g : Y convert H₂.comp (snd_embedding_of_left_embedding H₁ g) erw [← coe_comp] congr - rw [←limit.w _ WalkingCospan.Hom.inr] + rw [← limit.w _ WalkingCospan.Hom.inr] rfl #align Top.embedding_of_pullback_embeddings TopCat.embedding_of_pullback_embeddings @@ -364,7 +364,7 @@ theorem openEmbedding_of_pullback_open_embeddings {X Y S : TopCat} {f : X ⟶ S} convert H₂.comp (snd_openEmbedding_of_left_openEmbedding H₁ g) erw [← coe_comp] congr - rw [←(limit.w _ WalkingCospan.Hom.inr)] + rw [← limit.w _ WalkingCospan.Hom.inr] rfl #align Top.open_embedding_of_pullback_open_embeddings TopCat.openEmbedding_of_pullback_open_embeddings diff --git a/Mathlib/Topology/Clopen.lean b/Mathlib/Topology/Clopen.lean index 0e2dd4e2570db..2ca7a7ae83c13 100644 --- a/Mathlib/Topology/Clopen.lean +++ b/Mathlib/Topology/Clopen.lean @@ -112,7 +112,7 @@ theorem IsClopen.preimage {s : Set Y} (h : IsClopen s) {f : X → Y} (hf : Conti theorem ContinuousOn.preimage_isClopen_of_isClopen {f : X → Y} {s : Set X} {t : Set Y} (hf : ContinuousOn f s) (hs : IsClopen s) (ht : IsClopen t) : IsClopen (s ∩ f ⁻¹' t) := - ⟨ContinuousOn.preimage_open_of_open hf hs.1 ht.1, + ⟨ContinuousOn.isOpen_inter_preimage hf hs.1 ht.1, ContinuousOn.preimage_isClosed_of_isClosed hf hs.2 ht.2⟩ #align continuous_on.preimage_clopen_of_clopen ContinuousOn.preimage_isClopen_of_isClopen diff --git a/Mathlib/Topology/Compactness/Compact.lean b/Mathlib/Topology/Compactness/Compact.lean index da71c59517a1d..5bcdfe33ff9af 100644 --- a/Mathlib/Topology/Compactness/Compact.lean +++ b/Mathlib/Topology/Compactness/Compact.lean @@ -486,7 +486,7 @@ theorem exists_subset_nhds_of_isCompact' [Nonempty ι] {V : ι → Set X} obtain ⟨W, hsubW, W_op, hWU⟩ := exists_open_set_nhds hU suffices : ∃ i, V i ⊆ W · exact this.imp fun i hi => hi.trans hWU - by_contra' H + by_contra! H replace H : ∀ i, (V i ∩ Wᶜ).Nonempty := fun i => Set.inter_compl_nonempty_iff.mpr (H i) have : (⋂ i, V i ∩ Wᶜ).Nonempty := by refine' @@ -798,7 +798,7 @@ theorem finite_of_compact_of_discrete [CompactSpace X] [DiscreteTopology X] : Fi theorem exists_nhds_ne_neBot (X : Type*) [TopologicalSpace X] [CompactSpace X] [Infinite X] : ∃ x : X, (𝓝[≠] x).NeBot := by - by_contra' H + by_contra! H simp_rw [not_neBot] at H haveI := discreteTopology_iff_nhds_ne.2 H exact Infinite.not_finite (finite_of_compact_of_discrete : Finite X) @@ -938,7 +938,7 @@ theorem IsCompact.finite (hs : IsCompact s) (hs' : DiscreteTopology s) : s.Finit theorem exists_nhds_ne_inf_principal_neBot (hs : IsCompact s) (hs' : s.Infinite) : ∃ x ∈ s, (𝓝[≠] x ⊓ 𝓟 s).NeBot := by - by_contra' H + by_contra! H simp_rw [not_neBot] at H exact hs' (hs.finite <| discreteTopology_subtype_iff.mpr H) #align exists_nhds_ne_inf_principal_ne_bot exists_nhds_ne_inf_principal_neBot diff --git a/Mathlib/Topology/Connected/Basic.lean b/Mathlib/Topology/Connected/Basic.lean index cd287cca63ace..6369a5f29a852 100644 --- a/Mathlib/Topology/Connected/Basic.lean +++ b/Mathlib/Topology/Connected/Basic.lean @@ -371,8 +371,8 @@ theorem Inducing.isPreconnected_image [TopologicalSpace β] {s : Set α} {f : α /- TODO: The following lemmas about connection of preimages hold more generally for strict maps (the quotient and subspace topologies of the image agree) whose fibers are preconnected. -/ -theorem IsPreconnected.preimage_of_open_map [TopologicalSpace β] {s : Set β} (hs : IsPreconnected s) - {f : α → β} (hinj : Function.Injective f) (hf : IsOpenMap f) (hsf : s ⊆ range f) : +theorem IsPreconnected.preimage_of_isOpenMap [TopologicalSpace β] {f : α → β} {s : Set β} + (hs : IsPreconnected s) (hinj : Function.Injective f) (hf : IsOpenMap f) (hsf : s ⊆ range f) : IsPreconnected (f ⁻¹' s) := fun u v hu hv hsuv hsu hsv => by replace hsf : f '' (f ⁻¹' s) = s := image_preimage_eq_of_subset hsf obtain ⟨_, has, ⟨a, hau, rfl⟩, hav⟩ : (s ∩ (f '' u ∩ f '' v)).Nonempty @@ -381,7 +381,7 @@ theorem IsPreconnected.preimage_of_open_map [TopologicalSpace β] {s : Set β} ( · simpa only [image_preimage_inter] using hsu.image f · simpa only [image_preimage_inter] using hsv.image f · exact ⟨a, has, hau, hinj.mem_set_image.1 hav⟩ -#align is_preconnected.preimage_of_open_map IsPreconnected.preimage_of_open_map +#align is_preconnected.preimage_of_open_map IsPreconnected.preimage_of_isOpenMap theorem IsPreconnected.preimage_of_isClosedMap [TopologicalSpace β] {s : Set β} (hs : IsPreconnected s) {f : α → β} (hinj : Function.Injective f) (hf : IsClosedMap f) @@ -396,11 +396,11 @@ theorem IsPreconnected.preimage_of_isClosedMap [TopologicalSpace β] {s : Set β · exact ⟨a, has, hau, hinj.mem_set_image.1 hav⟩ #align is_preconnected.preimage_of_closed_map IsPreconnected.preimage_of_isClosedMap -theorem IsConnected.preimage_of_openMap [TopologicalSpace β] {s : Set β} (hs : IsConnected s) +theorem IsConnected.preimage_of_isOpenMap [TopologicalSpace β] {s : Set β} (hs : IsConnected s) {f : α → β} (hinj : Function.Injective f) (hf : IsOpenMap f) (hsf : s ⊆ range f) : IsConnected (f ⁻¹' s) := - ⟨hs.nonempty.preimage' hsf, hs.isPreconnected.preimage_of_open_map hinj hf hsf⟩ -#align is_connected.preimage_of_open_map IsConnected.preimage_of_openMap + ⟨hs.nonempty.preimage' hsf, hs.isPreconnected.preimage_of_isOpenMap hinj hf hsf⟩ +#align is_connected.preimage_of_open_map IsConnected.preimage_of_isOpenMap theorem IsConnected.preimage_of_isClosedMap [TopologicalSpace β] {s : Set β} (hs : IsConnected s) {f : α → β} (hinj : Function.Injective f) (hf : IsClosedMap f) (hsf : s ⊆ range f) : @@ -512,7 +512,7 @@ theorem Sigma.isConnected_iff [∀ i, TopologicalSpace (π i)] {s : Set (Σi, π · obtain ⟨⟨i, x⟩, hx⟩ := hs.nonempty have : s ⊆ range (Sigma.mk i) := hs.isPreconnected.subset_isClopen isClopen_range_sigmaMk ⟨⟨i, x⟩, hx, x, rfl⟩ - exact ⟨i, Sigma.mk i ⁻¹' s, hs.preimage_of_openMap sigma_mk_injective isOpenMap_sigmaMk this, + exact ⟨i, Sigma.mk i ⁻¹' s, hs.preimage_of_isOpenMap sigma_mk_injective isOpenMap_sigmaMk this, (Set.image_preimage_eq_of_subset this).symm⟩ · rintro ⟨i, t, ht, rfl⟩ exact ht.image _ continuous_sigmaMk.continuousOn @@ -537,12 +537,12 @@ theorem Sum.isConnected_iff [TopologicalSpace β] {s : Set (Sum α β)} : · have h : s ⊆ range Sum.inl := hs.isPreconnected.subset_isClopen isClopen_range_inl ⟨.inl x, hx, x, rfl⟩ refine' Or.inl ⟨Sum.inl ⁻¹' s, _, _⟩ - · exact hs.preimage_of_openMap Sum.inl_injective isOpenMap_inl h + · exact hs.preimage_of_isOpenMap Sum.inl_injective isOpenMap_inl h · exact (image_preimage_eq_of_subset h).symm · have h : s ⊆ range Sum.inr := hs.isPreconnected.subset_isClopen isClopen_range_inr ⟨.inr x, hx, x, rfl⟩ refine' Or.inr ⟨Sum.inr ⁻¹' s, _, _⟩ - · exact hs.preimage_of_openMap Sum.inr_injective isOpenMap_inr h + · exact hs.preimage_of_isOpenMap Sum.inr_injective isOpenMap_inr h · exact (image_preimage_eq_of_subset h).symm · rintro (⟨t, ht, rfl⟩ | ⟨t, ht, rfl⟩) · exact ht.image _ continuous_inl.continuousOn @@ -1311,7 +1311,7 @@ continuous on a set `s`, is constant on s, then s is preconnected -/ theorem isPreconnected_of_forall_constant {s : Set α} (hs : ∀ f : α → Bool, ContinuousOn f s → ∀ x ∈ s, ∀ y ∈ s, f x = f y) : IsPreconnected s := by unfold IsPreconnected - by_contra' + by_contra! rcases this with ⟨u, v, u_op, v_op, hsuv, ⟨x, x_in_s, x_in_u⟩, ⟨y, y_in_s, y_in_v⟩, H⟩ have hy : y ∉ u := fun y_in_u => eq_empty_iff_forall_not_mem.mp H y ⟨y_in_s, ⟨y_in_u, y_in_v⟩⟩ have : ContinuousOn u.boolIndicator s := by diff --git a/Mathlib/Topology/Connected/LocallyConnected.lean b/Mathlib/Topology/Connected/LocallyConnected.lean index 40f3ce3ad7111..8f7a3fcac7ae8 100644 --- a/Mathlib/Topology/Connected/LocallyConnected.lean +++ b/Mathlib/Topology/Connected/LocallyConnected.lean @@ -136,7 +136,7 @@ theorem OpenEmbedding.locallyConnectedSpace [LocallyConnectedSpace α] [Topologi {f : β → α} (h : OpenEmbedding f) : LocallyConnectedSpace β := by refine locallyConnectedSpace_of_connected_bases (fun _ s ↦ f ⁻¹' s) (fun x s ↦ (IsOpen s ∧ f x ∈ s ∧ IsConnected s) ∧ s ⊆ range f) (fun x ↦ ?_) - (fun x s hxs ↦ hxs.1.2.2.isPreconnected.preimage_of_open_map h.inj h.isOpenMap hxs.2) + (fun x s hxs ↦ hxs.1.2.2.isPreconnected.preimage_of_isOpenMap h.inj h.isOpenMap hxs.2) rw [h.nhds_eq_comap] exact LocallyConnectedSpace.open_connected_basis (f x) |>.restrict_subset (h.open_range.mem_nhds <| mem_range_self _) |>.comap _ diff --git a/Mathlib/Topology/Connected/PathConnected.lean b/Mathlib/Topology/Connected/PathConnected.lean index 46afc4cd61404..8b64b781f4113 100644 --- a/Mathlib/Topology/Connected/PathConnected.lean +++ b/Mathlib/Topology/Connected/PathConnected.lean @@ -172,12 +172,15 @@ def symm (γ : Path x y) : Path y x where #align path.symm Path.symm @[simp] -theorem symm_symm {γ : Path x y} : γ.symm.symm = γ := by +theorem symm_symm (γ : Path x y) : γ.symm.symm = γ := by ext t show γ (σ (σ t)) = γ t rw [unitInterval.symm_symm] #align path.symm_symm Path.symm_symm +theorem symm_bijective : Function.Bijective (Path.symm : Path x y → Path y x) := + Function.bijective_iff_has_inverse.mpr ⟨_, symm_symm, symm_symm⟩ + @[simp] theorem refl_symm {a : X} : (Path.refl a).symm = Path.refl a := by ext diff --git a/Mathlib/Topology/Connected/TotallyDisconnected.lean b/Mathlib/Topology/Connected/TotallyDisconnected.lean index 252995bd769b3..850230f42c47c 100644 --- a/Mathlib/Topology/Connected/TotallyDisconnected.lean +++ b/Mathlib/Topology/Connected/TotallyDisconnected.lean @@ -95,7 +95,7 @@ theorem isTotallyDisconnected_of_isClopen_set {X : Type*} [TopologicalSpace X] IsTotallyDisconnected (Set.univ : Set X) := by rintro S - hS unfold Set.Subsingleton - by_contra' h_contra + by_contra! h_contra rcases h_contra with ⟨x, hx, y, hy, hxy⟩ obtain ⟨U, hU, hxU, hyU⟩ := hX hxy specialize diff --git a/Mathlib/Topology/ContinuousFunction/Algebra.lean b/Mathlib/Topology/ContinuousFunction/Algebra.lean index 810ea1174b5d1..84639edf7dd3a 100644 --- a/Mathlib/Topology/ContinuousFunction/Algebra.lean +++ b/Mathlib/Topology/ContinuousFunction/Algebra.lean @@ -385,7 +385,7 @@ open BigOperators @[to_additive (attr := simp)] theorem coe_prod [CommMonoid β] [ContinuousMul β] {ι : Type*} (s : Finset ι) (f : ι → C(α, β)) : ⇑(∏ i in s, f i) = ∏ i in s, (f i : α → β) := - (coeFnMonoidHom : C(α, β) →* _).map_prod f s + map_prod coeFnMonoidHom f s #align continuous_map.coe_prod ContinuousMap.coe_prod #align continuous_map.coe_sum ContinuousMap.coe_sum @@ -1042,10 +1042,10 @@ theorem periodic_tsum_comp_add_zsmul [AddCommGroup X] [TopologicalAddGroup X] [A by_cases h : Summable fun n : ℤ => f.comp (ContinuousMap.addRight (n • p)) · convert congr_arg (fun f : C(X, Y) => f x) ((Equiv.addRight (1 : ℤ)).tsum_eq _) using 1 -- Porting note: in mathlib3 the proof from here was: - -- simp_rw [←tsum_apply h, ←tsum_apply ((equiv.add_right (1 : ℤ)).summable_iff.mpr h), + -- simp_rw [← tsum_apply h, ← tsum_apply ((equiv.add_right (1 : ℤ)).summable_iff.mpr h), -- equiv.coe_add_right, comp_apply, coe_add_right, add_one_zsmul, add_comm (_ • p) p, - -- ←add_assoc] - -- However now the second `←tsum_apply` doesn't fire unless we use `erw`. + -- ← add_assoc] + -- However now the second `← tsum_apply` doesn't fire unless we use `erw`. simp_rw [← tsum_apply h] erw [← tsum_apply ((Equiv.addRight (1 : ℤ)).summable_iff.mpr h)] simp [coe_addRight, add_one_zsmul, add_comm (_ • p) p, ← add_assoc] diff --git a/Mathlib/Topology/ContinuousFunction/Bounded.lean b/Mathlib/Topology/ContinuousFunction/Bounded.lean index bc4394142c5d4..6679340844d64 100644 --- a/Mathlib/Topology/ContinuousFunction/Bounded.lean +++ b/Mathlib/Topology/ContinuousFunction/Bounded.lean @@ -767,7 +767,7 @@ open BigOperators @[simp] theorem coe_sum {ι : Type*} (s : Finset ι) (f : ι → α →ᵇ β) : ⇑(∑ i in s, f i) = ∑ i in s, (f i : α → β) := - (@coeFnAddHom α β _ _ _ _).map_sum f s + map_sum coeFnAddHom f s #align bounded_continuous_function.coe_sum BoundedContinuousFunction.coe_sum theorem sum_apply {ι : Type*} (s : Finset ι) (f : ι → α →ᵇ β) (a : α) : diff --git a/Mathlib/Topology/ContinuousOn.lean b/Mathlib/Topology/ContinuousOn.lean index b576ff2a120ac..a35725ca6dba1 100644 --- a/Mathlib/Topology/ContinuousOn.lean +++ b/Mathlib/Topology/ContinuousOn.lean @@ -1064,10 +1064,10 @@ theorem continuousOn_open_iff {f : α → β} {s : Set α} (hs : IsOpen s) : rw [@inter_comm _ s (f ⁻¹' t), inter_assoc, inter_self] #align continuous_on_open_iff continuousOn_open_iff -theorem ContinuousOn.preimage_open_of_open {f : α → β} {s : Set α} {t : Set β} +theorem ContinuousOn.isOpen_inter_preimage {f : α → β} {s : Set α} {t : Set β} (hf : ContinuousOn f s) (hs : IsOpen s) (ht : IsOpen t) : IsOpen (s ∩ f ⁻¹' t) := (continuousOn_open_iff hs).1 hf t ht -#align continuous_on.preimage_open_of_open ContinuousOn.preimage_open_of_open +#align continuous_on.preimage_open_of_open ContinuousOn.isOpen_inter_preimage theorem ContinuousOn.isOpen_preimage {f : α → β} {s : Set α} {t : Set β} (h : ContinuousOn f s) (hs : IsOpen s) (hp : f ⁻¹' t ⊆ s) (ht : IsOpen t) : IsOpen (f ⁻¹' t) := by @@ -1087,7 +1087,7 @@ theorem ContinuousOn.preimage_interior_subset_interior_preimage {f : α → β} calc s ∩ f ⁻¹' interior t ⊆ interior (s ∩ f ⁻¹' t) := interior_maximal (inter_subset_inter (Subset.refl _) (preimage_mono interior_subset)) - (hf.preimage_open_of_open hs isOpen_interior) + (hf.isOpen_inter_preimage hs isOpen_interior) _ = s ∩ interior (f ⁻¹' t) := by rw [interior_inter, hs.interior_eq] #align continuous_on.preimage_interior_subset_interior_preimage ContinuousOn.preimage_interior_subset_interior_preimage @@ -1109,12 +1109,12 @@ theorem continuousOn_to_generateFrom_iff {s : Set α} {T : Set (Set β)} {f : α exact forall_congr' fun t => forall_swap -- porting note: dropped an unneeded assumption -theorem continuousOn_open_of_generateFrom {β : Type*} {s : Set α} {T : Set (Set β)} {f : α → β} +theorem continuousOn_isOpen_of_generateFrom {β : Type*} {s : Set α} {T : Set (Set β)} {f : α → β} (h : ∀ t ∈ T, IsOpen (s ∩ f ⁻¹' t)) : @ContinuousOn α β _ (.generateFrom T) f s := continuousOn_to_generateFrom_iff.2 fun _x hx t ht hxt => mem_nhdsWithin.2 ⟨_, h t ht, ⟨hx, hxt⟩, fun _y hy => hy.1.2⟩ -#align continuous_on_open_of_generate_from continuousOn_open_of_generateFromₓ +#align continuous_on_open_of_generate_from continuousOn_isOpen_of_generateFromₓ theorem ContinuousWithinAt.prod {f : α → β} {g : α → γ} {s : Set α} {x : α} (hf : ContinuousWithinAt f s x) (hg : ContinuousWithinAt g s x) : diff --git a/Mathlib/Topology/Covering.lean b/Mathlib/Topology/Covering.lean index b5280cf15b8c6..971ad4b7ab7a1 100644 --- a/Mathlib/Topology/Covering.lean +++ b/Mathlib/Topology/Covering.lean @@ -181,7 +181,7 @@ protected theorem isSeparatedMap : IsSeparatedMap f := refine ⟨t.source ∩ (Prod.snd ∘ t) ⁻¹' {(t e₁).2}, t.source ∩ (Prod.snd ∘ t) ⁻¹' {(t e₂).2}, ?_, ?_, ⟨he₁, rfl⟩, ⟨he₂, rfl⟩, Set.disjoint_left.mpr fun x h₁ h₂ ↦ hne (t.injOn he₁ he₂ ?_)⟩ iterate 2 - exact t.continuous_toFun.preimage_open_of_open t.open_source + exact t.continuous_toFun.isOpen_inter_preimage t.open_source (continuous_snd.isOpen_preimage _ <| isOpen_discrete _) refine Prod.ext ?_ (h₁.2.symm.trans h₂.2) rwa [t.proj_toFun e₁ he₁, t.proj_toFun e₂ he₂] diff --git a/Mathlib/Topology/DiscreteSubset.lean b/Mathlib/Topology/DiscreteSubset.lean index 9f2c503bbd1eb..447b8d33c1ade 100644 --- a/Mathlib/Topology/DiscreteSubset.lean +++ b/Mathlib/Topology/DiscreteSubset.lean @@ -84,7 +84,7 @@ theorem isClosed_and_discrete_iff {S : Set X} : IsClosed S ∧ DiscreteTopology S ↔ ∀ x, Disjoint (𝓝[≠] x) (𝓟 S) := by rw [discreteTopology_subtype_iff, isClosed_iff_clusterPt, ← forall_and] congrm (∀ x, ?_) - rw [← not_imp_not, clusterPt_iff_not_disjoint, not_not, ←disjoint_iff] + rw [← not_imp_not, clusterPt_iff_not_disjoint, not_not, ← disjoint_iff] constructor <;> intro H · by_cases hx : x ∈ S exacts [H.2 hx, (H.1 hx).mono_left nhdsWithin_le_nhds] @@ -97,11 +97,11 @@ def Filter.codiscrete (X : Type*) [TopologicalSpace X] : Filter X where univ_sets := ⟨isOpen_univ, compl_univ.symm ▸ Subsingleton.discreteTopology⟩ sets_of_superset := by intro U V hU hV - simp_rw [←isClosed_compl_iff, isClosed_and_discrete_iff] at hU ⊢ + simp_rw [← isClosed_compl_iff, isClosed_and_discrete_iff] at hU ⊢ exact fun x ↦ (hU x).mono_right (principal_mono.mpr <| compl_subset_compl.mpr hV) inter_sets := by intro U V hU hV - simp_rw [←isClosed_compl_iff, isClosed_and_discrete_iff] at hU hV ⊢ + simp_rw [← isClosed_compl_iff, isClosed_and_discrete_iff] at hU hV ⊢ exact fun x ↦ compl_inter U V ▸ sup_principal ▸ disjoint_sup_right.mpr ⟨hU x, hV x⟩ end codiscrete_filter diff --git a/Mathlib/Topology/EMetricSpace/Paracompact.lean b/Mathlib/Topology/EMetricSpace/Paracompact.lean index b4494c9f093de..ba1e0e2ac41ea 100644 --- a/Mathlib/Topology/EMetricSpace/Paracompact.lean +++ b/Mathlib/Topology/EMetricSpace/Paracompact.lean @@ -91,7 +91,7 @@ instance (priority := 100) instParacompactSpace [PseudoEMetricSpace α] : Paraco rcases ENNReal.exists_inv_two_pow_lt this.ne' with ⟨n, hn⟩ refine' ⟨n, Subset.trans (ball_subset_ball _) hε⟩ simpa only [div_eq_mul_inv, mul_comm] using (ENNReal.mul_lt_of_lt_div hn).le - by_contra' h + by_contra! h apply h n (ind x) exact memD.2 ⟨x, rfl, hn, fun _ _ _ => h _ _, mem_ball_self (pow_pos _)⟩ -- Each `D n i` is a union of open balls, hence it is an open set @@ -138,7 +138,7 @@ instance (priority := 100) instParacompactSpace [PseudoEMetricSpace α] : Paraco -- For each `m ≤ n + k` there is at most one `j` such that `D m j ∩ B` is nonempty. have Hle : ∀ m ≤ n + k, Set.Subsingleton { j | (D m j ∩ B).Nonempty } := by rintro m hm j₁ ⟨y, hyD, hyB⟩ j₂ ⟨z, hzD, hzB⟩ - by_contra' h' : j₁ ≠ j₂ + by_contra! h' : j₁ ≠ j₂ wlog h : j₁ < j₂ generalizing j₁ j₂ y z · exact this z hzD hzB y hyD hyB h'.symm (h'.lt_or_lt.resolve_left h) rcases memD.1 hyD with ⟨y', rfl, hsuby, -, hdisty⟩ diff --git a/Mathlib/Topology/FiberBundle/Basic.lean b/Mathlib/Topology/FiberBundle/Basic.lean index cc72feb8c05ec..d50518b3d983a 100644 --- a/Mathlib/Topology/FiberBundle/Basic.lean +++ b/Mathlib/Topology/FiberBundle/Basic.lean @@ -579,7 +579,7 @@ def localTriv (i : ι) : Trivialization F Z.proj where simp only [exists_prop, mem_iUnion, mem_singleton_iff] exact ⟨i, s, s_open, rfl⟩ continuous_invFun := by - refine continuousOn_open_of_generateFrom fun t ht ↦ ?_ + refine continuousOn_isOpen_of_generateFrom fun t ht ↦ ?_ simp only [exists_prop, mem_iUnion, mem_singleton_iff] at ht obtain ⟨j, s, s_open, ts⟩ : ∃ j s, IsOpen s ∧ t = (localTrivAsLocalEquiv Z j).source ∩ localTrivAsLocalEquiv Z j ⁻¹' s := ht diff --git a/Mathlib/Topology/GDelta.lean b/Mathlib/Topology/GDelta.lean index 5328101d8d878..d48ade9b5fabb 100644 --- a/Mathlib/Topology/GDelta.lean +++ b/Mathlib/Topology/GDelta.lean @@ -74,16 +74,16 @@ theorem isGδ_univ : IsGδ (univ : Set X) := isOpen_univ.isGδ #align is_Gδ_univ isGδ_univ -theorem isGδ_biInter_of_open {I : Set ι} (hI : I.Countable) {f : ι → Set X} +theorem isGδ_biInter_of_isOpen {I : Set ι} (hI : I.Countable) {f : ι → Set X} (hf : ∀ i ∈ I, IsOpen (f i)) : IsGδ (⋂ i ∈ I, f i) := ⟨f '' I, by rwa [ball_image_iff], hI.image _, by rw [sInter_image]⟩ -#align is_Gδ_bInter_of_open isGδ_biInter_of_open +#align is_Gδ_bInter_of_open isGδ_biInter_of_isOpen -- porting note: TODO: generalize to `Sort*` + `Countable _` -theorem isGδ_iInter_of_open [Encodable ι] {f : ι → Set X} (hf : ∀ i, IsOpen (f i)) : +theorem isGδ_iInter_of_isOpen [Encodable ι] {f : ι → Set X} (hf : ∀ i, IsOpen (f i)) : IsGδ (⋂ i, f i) := ⟨range f, by rwa [forall_range_iff], countable_range _, by rw [sInter_range]⟩ -#align is_Gδ_Inter_of_open isGδ_iInter_of_open +#align is_Gδ_Inter_of_open isGδ_iInter_of_isOpen -- porting note: TODO: generalize to `Sort*` + `Countable _` /-- The intersection of an encodable family of Gδ sets is a Gδ set. -/ @@ -116,7 +116,7 @@ theorem IsGδ.union {s t : Set X} (hs : IsGδ s) (ht : IsGδ t) : IsGδ (s ∪ t rcases hs with ⟨S, Sopen, Scount, rfl⟩ rcases ht with ⟨T, Topen, Tcount, rfl⟩ rw [sInter_union_sInter] - apply isGδ_biInter_of_open (Scount.prod Tcount) + apply isGδ_biInter_of_isOpen (Scount.prod Tcount) rintro ⟨a, b⟩ ⟨ha, hb⟩ exact (Sopen a ha).union (Topen b hb) #align is_Gδ.union IsGδ.union diff --git a/Mathlib/Topology/Gluing.lean b/Mathlib/Topology/Gluing.lean index a970c8ea77e8c..9bdf29f93691a 100644 --- a/Mathlib/Topology/Gluing.lean +++ b/Mathlib/Topology/Gluing.lean @@ -211,14 +211,14 @@ theorem ι_eq_iff_rel (i j : D.J) (x : D.U i) (y : D.U j) : rw [← (InvImage.equivalence _ _ D.rel_equiv).eqvGen_iff] refine' EqvGen.mono _ (D.eqvGen_of_π_eq h : _) rintro _ _ ⟨x⟩ - rw [←show (sigmaIsoSigma.{u, u} _).inv _ = x from + rw [← show (sigmaIsoSigma.{u, u} _).inv _ = x from ConcreteCategory.congr_hom (sigmaIsoSigma.{u, u} _).hom_inv_id x] generalize (sigmaIsoSigma.{u, u} D.V).hom x = x' obtain ⟨⟨i, j⟩, y⟩ := x' unfold InvImage MultispanIndex.fstSigmaMap MultispanIndex.sndSigmaMap simp only [Opens.inclusion_apply, TopCat.comp_app, sigmaIsoSigma_inv_apply, Cofan.mk_ι_app] - rw [←comp_apply, colimit.ι_desc, ←comp_apply, colimit.ι_desc] + rw [← comp_apply, colimit.ι_desc, ← comp_apply, colimit.ι_desc] erw [sigmaIsoSigma_hom_ι_apply, sigmaIsoSigma_hom_ι_apply] exact Or.inr ⟨y, ⟨rfl, rfl⟩⟩ · rintro (⟨⟨⟩⟩ | ⟨z, e₁, e₂⟩) diff --git a/Mathlib/Topology/Homeomorph.lean b/Mathlib/Topology/Homeomorph.lean index 926f7bfa3f155..54a38c1c0cbfa 100644 --- a/Mathlib/Topology/Homeomorph.lean +++ b/Mathlib/Topology/Homeomorph.lean @@ -79,6 +79,9 @@ protected def symm (h : X ≃ₜ Y) : Y ≃ₜ X where @[simp] theorem symm_symm (h : X ≃ₜ Y) : h.symm.symm = h := rfl #align homeomorph.symm_symm Homeomorph.symm_symm +theorem symm_bijective : Function.Bijective (Homeomorph.symm : (X ≃ₜ Y) → Y ≃ₜ X) := + Function.bijective_iff_has_inverse.mpr ⟨_, symm_symm, symm_symm⟩ + /-- See Note [custom simps projection] -/ def Simps.symm_apply (h : X ≃ₜ Y) : Y → X := h.symm diff --git a/Mathlib/Topology/Homotopy/Basic.lean b/Mathlib/Topology/Homotopy/Basic.lean index 430eec96f32d6..7278cec8eea98 100644 --- a/Mathlib/Topology/Homotopy/Basic.lean +++ b/Mathlib/Topology/Homotopy/Basic.lean @@ -231,6 +231,10 @@ theorem symm_symm {f₀ f₁ : C(X, Y)} (F : Homotopy f₀ f₁) : F.symm.symm = simp #align continuous_map.homotopy.symm_symm ContinuousMap.Homotopy.symm_symm +theorem symm_bijective {f₀ f₁ : C(X, Y)} : + Function.Bijective (Homotopy.symm : Homotopy f₀ f₁ → Homotopy f₁ f₀) := + Function.bijective_iff_has_inverse.mpr ⟨_, symm_symm, symm_symm⟩ + /-- Given `Homotopy f₀ f₁` and `Homotopy f₁ f₂`, we can define a `Homotopy f₀ f₂` by putting the first homotopy on `[0, 1/2]` and the second on `[1/2, 1]`. @@ -503,6 +507,10 @@ theorem symm_symm {f₀ f₁ : C(X, Y)} (F : HomotopyWith f₀ f₁ P) : F.symm. ext <| Homotopy.congr_fun <| Homotopy.symm_symm _ #align continuous_map.homotopy_with.symm_symm ContinuousMap.HomotopyWith.symm_symm +theorem symm_bijective {f₀ f₁ : C(X, Y)} : + Function.Bijective (HomotopyWith.symm : HomotopyWith f₀ f₁ P → HomotopyWith f₁ f₀ P) := + Function.bijective_iff_has_inverse.mpr ⟨_, symm_symm, symm_symm⟩ + /-- Given `HomotopyWith f₀ f₁ P` and `HomotopyWith f₁ f₂ P`, we can define a `HomotopyWith f₀ f₂ P` by putting the first homotopy on `[0, 1/2]` and the second on `[1/2, 1]`. @@ -629,6 +637,10 @@ theorem symm_symm (F : HomotopyRel f₀ f₁ S) : F.symm.symm = F := HomotopyWith.symm_symm F #align continuous_map.homotopy_rel.symm_symm ContinuousMap.HomotopyRel.symm_symm +theorem symm_bijective : + Function.Bijective (HomotopyRel.symm : HomotopyRel f₀ f₁ S → HomotopyRel f₁ f₀ S) := + Function.bijective_iff_has_inverse.mpr ⟨_, symm_symm, symm_symm⟩ + /-- Given `HomotopyRel f₀ f₁ S` and `HomotopyRel f₁ f₂ S`, we can define a `HomotopyRel f₀ f₂ S` by putting the first homotopy on `[0, 1/2]` and the second on `[1/2, 1]`. -/ diff --git a/Mathlib/Topology/Homotopy/Path.lean b/Mathlib/Topology/Homotopy/Path.lean index df37daa49a202..dc7b288f306c2 100644 --- a/Mathlib/Topology/Homotopy/Path.lean +++ b/Mathlib/Topology/Homotopy/Path.lean @@ -117,6 +117,9 @@ theorem symm_symm (F : Homotopy p₀ p₁) : F.symm.symm = F := ContinuousMap.HomotopyRel.symm_symm F #align path.homotopy.symm_symm Path.Homotopy.symm_symm +theorem symm_bijective : Function.Bijective (Homotopy.symm : Homotopy p₀ p₁ → Homotopy p₁ p₀) := + Function.bijective_iff_has_inverse.mpr ⟨_, symm_symm, symm_symm⟩ + /-- Given `Homotopy p₀ p₁` and `Homotopy p₁ p₂`, we can define a `Homotopy p₀ p₂` by putting the first homotopy on `[0, 1/2]` and the second on `[1/2, 1]`. diff --git a/Mathlib/Topology/Instances/AddCircle.lean b/Mathlib/Topology/Instances/AddCircle.lean index d53de1054f606..bbd27374a501b 100644 --- a/Mathlib/Topology/Instances/AddCircle.lean +++ b/Mathlib/Topology/Instances/AddCircle.lean @@ -409,7 +409,7 @@ theorem addOrderOf_eq_pos_iff {u : AddCircle p} {n : ℕ} (h : 0 < n) : have he : (↑(↑((a % n).toNat) / ↑n * p) : AddCircle p) = k · convert congr_arg (QuotientAddGroup.mk : 𝕜 → (AddCircle p)) ha using 1 rw [coe_add, ← Int.cast_ofNat, han, zsmul_eq_mul, mul_div_right_comm, eq_comm, add_left_eq_self, - ←zsmul_eq_mul, coe_zsmul, coe_period, smul_zero] + ← zsmul_eq_mul, coe_zsmul, coe_period, smul_zero] refine' ⟨(a % n).toNat, _, _, he⟩ · rw [← Int.ofNat_lt, han] exact Int.emod_lt_of_pos _ (Int.ofNat_lt.2 h) diff --git a/Mathlib/Topology/Instances/RatLemmas.lean b/Mathlib/Topology/Instances/RatLemmas.lean index 8456bd0812040..139c858ac8be5 100644 --- a/Mathlib/Topology/Instances/RatLemmas.lean +++ b/Mathlib/Topology/Instances/RatLemmas.lean @@ -81,7 +81,7 @@ theorem not_secondCountableTopology_opc : ¬SecondCountableTopology ℚ∞ := by instance : TotallyDisconnectedSpace ℚ := by refine' ⟨fun s hsu hs x hx y hy => _⟩; clear hsu - by_contra' H : x ≠ y + by_contra! H : x ≠ y wlog hlt : x < y · refine' this s hs y hy x hx H.symm <| H.lt_or_lt.resolve_left hlt <;> assumption rcases exists_irrational_btwn (Rat.cast_lt.2 hlt) with ⟨z, hz, hxz, hzy⟩ diff --git a/Mathlib/Topology/Instances/Real.lean b/Mathlib/Topology/Instances/Real.lean index 2a6dd6d4b6d8c..682b3c8d69133 100644 --- a/Mathlib/Topology/Instances/Real.lean +++ b/Mathlib/Topology/Instances/Real.lean @@ -59,7 +59,7 @@ instance : SecondCountableTopology ℝ := secondCountable_of_proper theorem Real.isTopologicalBasis_Ioo_rat : @IsTopologicalBasis ℝ _ (⋃ (a : ℚ) (b : ℚ) (_ : a < b), {Ioo (a : ℝ) b}) := - isTopologicalBasis_of_open_of_nhds (by simp (config := { contextual := true }) [isOpen_Ioo]) + isTopologicalBasis_of_isOpen_of_nhds (by simp (config := { contextual := true }) [isOpen_Ioo]) fun a v hav hv => let ⟨l, u, ⟨hl, hu⟩, h⟩ := mem_nhds_iff_exists_Ioo_subset.mp (IsOpen.mem_nhds hv hav) let ⟨q, hlq, hqa⟩ := exists_rat_btwn hl @@ -230,7 +230,7 @@ instance {a : ℝ} : DiscreteTopology (AddSubgroup.zmultiples a) := by rcases eq_or_ne a 0 with (rfl | ha) · rw [AddSubgroup.zmultiples_zero_eq_bot] exact Subsingleton.discreteTopology (α := (⊥ : Submodule ℤ ℝ)) - rw [discreteTopology_iff_open_singleton_zero, isOpen_induced_iff] + rw [discreteTopology_iff_isOpen_singleton_zero, isOpen_induced_iff] refine' ⟨ball 0 |a|, isOpen_ball, _⟩ ext ⟨x, hx⟩ obtain ⟨k, rfl⟩ := AddSubgroup.mem_zmultiples_iff.mp hx diff --git a/Mathlib/Topology/IsLocallyHomeomorph.lean b/Mathlib/Topology/IsLocallyHomeomorph.lean index 719db40c2d10b..e39f742bf10c5 100644 --- a/Mathlib/Topology/IsLocallyHomeomorph.lean +++ b/Mathlib/Topology/IsLocallyHomeomorph.lean @@ -203,12 +203,12 @@ open TopologicalSpace in /-- Ranges of continuous local sections of a local homeomorphism form a basis of the source space.-/ theorem isTopologicalBasis (hf : IsLocallyHomeomorph f) : IsTopologicalBasis {U : Set X | ∃ V : Set Y, IsOpen V ∧ ∃ s : C(V,X), f ∘ s = (↑) ∧ Set.range s = U} := by - refine isTopologicalBasis_of_open_of_nhds ?_ fun x U hx hU ↦ ?_ + refine isTopologicalBasis_of_isOpen_of_nhds ?_ fun x U hx hU ↦ ?_ · rintro _ ⟨U, hU, s, hs, rfl⟩ refine (openEmbedding_of_comp hf (hs ▸ ⟨embedding_subtype_val, ?_⟩) s.continuous).open_range rwa [Subtype.range_val] · obtain ⟨f, hxf, rfl⟩ := hf x - refine ⟨f.source ∩ U, ⟨f.target ∩ f.symm ⁻¹' U, f.symm.preimage_open_of_open hU, + refine ⟨f.source ∩ U, ⟨f.target ∩ f.symm ⁻¹' U, f.symm.isOpen_inter_preimage hU, ⟨_, continuousOn_iff_continuous_restrict.mp (f.continuous_invFun.mono fun _ h ↦ h.1)⟩, ?_, (Set.range_restrict _ _).trans ?_⟩, ⟨hxf, hx⟩, fun _ h ↦ h.2⟩ · ext y; exact f.right_inv y.2.1 diff --git a/Mathlib/Topology/LocalHomeomorph.lean b/Mathlib/Topology/LocalHomeomorph.lean index 8f28c723892dd..3978638bef564 100644 --- a/Mathlib/Topology/LocalHomeomorph.lean +++ b/Mathlib/Topology/LocalHomeomorph.lean @@ -356,6 +356,10 @@ theorem symm_target : e.symm.target = e.source := @[simp, mfld_simps] theorem symm_symm : e.symm.symm = e := rfl #align local_homeomorph.symm_symm LocalHomeomorph.symm_symm +theorem symm_bijective : Function.Bijective + (LocalHomeomorph.symm : LocalHomeomorph α β → LocalHomeomorph β α) := + Function.bijective_iff_has_inverse.mpr ⟨_, symm_symm, symm_symm⟩ + /-- A local homeomorphism is continuous at any point of its source -/ protected theorem continuousAt {x : α} (h : x ∈ e.source) : ContinuousAt e x := (e.continuousOn x h).continuousAt (e.open_source.mem_nhds h) @@ -441,15 +445,15 @@ theorem preimage_eventuallyEq_target_inter_preimage_inter {e : LocalHomeomorph e.left_inv hy, iff_true_intro hyu] #align local_homeomorph.preimage_eventually_eq_target_inter_preimage_inter LocalHomeomorph.preimage_eventuallyEq_target_inter_preimage_inter -theorem preimage_open_of_open {s : Set β} (hs : IsOpen s) : IsOpen (e.source ∩ e ⁻¹' s) := - e.continuousOn.preimage_open_of_open e.open_source hs -#align local_homeomorph.preimage_open_of_open LocalHomeomorph.preimage_open_of_open +theorem isOpen_inter_preimage {s : Set β} (hs : IsOpen s) : IsOpen (e.source ∩ e ⁻¹' s) := + e.continuousOn.isOpen_inter_preimage e.open_source hs +#align local_homeomorph.preimage_open_of_open LocalHomeomorph.isOpen_inter_preimage /-- A local homeomorphism is an open map on its source. -/ lemma isOpen_image_of_subset_source {s : Set α} (hs : IsOpen s) (hse : s ⊆ e.source) : IsOpen (e '' s) := by rw [(image_eq_target_inter_inv_preimage (e := e) hse)] - exact e.continuous_invFun.preimage_open_of_open e.open_target hs + exact e.continuous_invFun.isOpen_inter_preimage e.open_target hs /-- The inverse of a local homeomorphism `e` is an open map on `e.target`. -/ lemma isOpen_image_symm_of_subset_target {t : Set β} (ht : IsOpen t) (hte : t ⊆ e.target) : @@ -615,8 +619,8 @@ protected theorem frontier (h : e.IsImage s t) : e.IsImage (frontier s) (frontie #align local_homeomorph.is_image.frontier LocalHomeomorph.IsImage.frontier theorem isOpen_iff (h : e.IsImage s t) : IsOpen (e.source ∩ s) ↔ IsOpen (e.target ∩ t) := - ⟨fun hs => h.symm_preimage_eq' ▸ e.symm.preimage_open_of_open hs, fun hs => - h.preimage_eq' ▸ e.preimage_open_of_open hs⟩ + ⟨fun hs => h.symm_preimage_eq' ▸ e.symm.isOpen_inter_preimage hs, fun hs => + h.preimage_eq' ▸ e.isOpen_inter_preimage hs⟩ #align local_homeomorph.is_image.is_open_iff LocalHomeomorph.IsImage.isOpen_iff /-- Restrict a `LocalHomeomorph` to a pair of corresponding open sets. -/ @@ -657,21 +661,22 @@ theorem preimage_frontier (s : Set β) : (IsImage.of_preimage_eq rfl).frontier.preimage_eq #align local_homeomorph.preimage_frontier LocalHomeomorph.preimage_frontier -theorem preimage_open_of_open_symm {s : Set α} (hs : IsOpen s) : IsOpen (e.target ∩ e.symm ⁻¹' s) := - e.symm.continuousOn.preimage_open_of_open e.open_target hs -#align local_homeomorph.preimage_open_of_open_symm LocalHomeomorph.preimage_open_of_open_symm +theorem isOpen_inter_preimage_symm {s : Set α} (hs : IsOpen s) : IsOpen (e.target ∩ e.symm ⁻¹' s) := + e.symm.continuousOn.isOpen_inter_preimage e.open_target hs +#align local_homeomorph.preimage_open_of_open_symm LocalHomeomorph.isOpen_inter_preimage_symm /-- The image of an open set in the source is open. -/ -theorem image_open_of_open {s : Set α} (hs : IsOpen s) (h : s ⊆ e.source) : IsOpen (e '' s) := by +theorem image_isOpen_of_isOpen {s : Set α} (hs : IsOpen s) (h : s ⊆ e.source) : + IsOpen (e '' s) := by have : e '' s = e.target ∩ e.symm ⁻¹' s := e.toLocalEquiv.image_eq_target_inter_inv_preimage h rw [this] - exact e.continuousOn_symm.preimage_open_of_open e.open_target hs -#align local_homeomorph.image_open_of_open LocalHomeomorph.image_open_of_open + exact e.continuousOn_symm.isOpen_inter_preimage e.open_target hs +#align local_homeomorph.image_open_of_open LocalHomeomorph.image_isOpen_of_isOpen /-- The image of the restriction of an open set to the source is open. -/ -theorem image_open_of_open' {s : Set α} (hs : IsOpen s) : IsOpen (e '' (e.source ∩ s)) := - image_open_of_open _ (IsOpen.inter e.open_source hs) (inter_subset_left _ _) -#align local_homeomorph.image_open_of_open' LocalHomeomorph.image_open_of_open' +theorem image_isOpen_of_isOpen' {s : Set α} (hs : IsOpen s) : IsOpen (e '' (e.source ∩ s)) := + image_isOpen_of_isOpen _ (IsOpen.inter e.open_source hs) (inter_subset_left _ _) +#align local_homeomorph.image_open_of_open' LocalHomeomorph.image_isOpen_of_isOpen' /-- A `LocalEquiv` with continuous open forward map and an open source is a `LocalHomeomorph`. -/ def ofContinuousOpenRestrict (e : LocalEquiv α β) (hc : ContinuousOn e e.source) @@ -1285,7 +1290,8 @@ theorem openEmbedding_restrict : OpenEmbedding (e.source.restrict e) := by refine openEmbedding_of_continuous_injective_open (e.continuousOn.comp_continuous continuous_subtype_val Subtype.prop) e.injOn.injective fun V hV ↦ ?_ rw [Set.restrict_eq, Set.image_comp] - exact e.image_open_of_open (e.open_source.isOpenMap_subtype_val V hV) fun _ ⟨x, _, h⟩ ↦ h ▸ x.2 + exact e.image_isOpen_of_isOpen (e.open_source.isOpenMap_subtype_val V hV) + fun _ ⟨x, _, h⟩ ↦ h ▸ x.2 /-- A local homeomorphism whose source is all of `α` defines an open embedding of `α` into `β`. The converse is also true; see `OpenEmbedding.toLocalHomeomorph`. -/ @@ -1415,7 +1421,7 @@ theorem subtypeRestr_symm_trans_subtypeRestr (f f' : LocalHomeomorph α β) : (f.subtypeRestr s).symm.trans (f'.subtypeRestr s) ≈ (f.symm.trans f').restr (f.target ∩ f.symm ⁻¹' s) := by simp only [subtypeRestr_def, trans_symm_eq_symm_trans_symm] - have openness₁ : IsOpen (f.target ∩ f.symm ⁻¹' s) := f.preimage_open_of_open_symm s.2 + have openness₁ : IsOpen (f.target ∩ f.symm ⁻¹' s) := f.isOpen_inter_preimage_symm s.2 rw [← ofSet_trans _ openness₁, ← trans_assoc, ← trans_assoc] refine' EqOnSource.trans' _ (eqOnSource_refl _) -- f' has been eliminated !!! diff --git a/Mathlib/Topology/MetricSpace/Baire.lean b/Mathlib/Topology/MetricSpace/Baire.lean index 0cdf4279533ad..875a80922ec47 100644 --- a/Mathlib/Topology/MetricSpace/Baire.lean +++ b/Mathlib/Topology/MetricSpace/Baire.lean @@ -40,7 +40,7 @@ open EMetric ENNReal /-- The property `BaireSpace α` means that the topological space `α` has the Baire property: any countable intersection of open dense subsets is dense. -Formulated here when the source space is ℕ (and subsumed below by `dense_iInter_of_open` working +Formulated here when the source space is ℕ (and subsumed below by `dense_iInter_of_isOpen` working with any encodable source space). -/ class BaireSpace (α : Type*) [TopologicalSpace α] : Prop where baire_property : ∀ f : ℕ → Set α, (∀ n, IsOpen (f n)) → (∀ n, Dense (f n)) → Dense (⋂ n, f n) @@ -187,43 +187,43 @@ instance (priority := 100) BaireSpace.of_t2Space_locallyCompactSpace variable [TopologicalSpace α] [BaireSpace α] /-- Definition of a Baire space. -/ -theorem dense_iInter_of_open_nat {f : ℕ → Set α} (ho : ∀ n, IsOpen (f n)) (hd : ∀ n, Dense (f n)) : - Dense (⋂ n, f n) := +theorem dense_iInter_of_isOpen_nat {f : ℕ → Set α} (ho : ∀ n, IsOpen (f n)) + (hd : ∀ n, Dense (f n)) : Dense (⋂ n, f n) := BaireSpace.baire_property f ho hd -#align dense_Inter_of_open_nat dense_iInter_of_open_nat +#align dense_Inter_of_open_nat dense_iInter_of_isOpen_nat /-- Baire theorem: a countable intersection of dense open sets is dense. Formulated here with ⋂₀. -/ -theorem dense_sInter_of_open {S : Set (Set α)} (ho : ∀ s ∈ S, IsOpen s) (hS : S.Countable) +theorem dense_sInter_of_isOpen {S : Set (Set α)} (ho : ∀ s ∈ S, IsOpen s) (hS : S.Countable) (hd : ∀ s ∈ S, Dense s) : Dense (⋂₀ S) := by cases' S.eq_empty_or_nonempty with h h · simp [h] · rcases hS.exists_eq_range h with ⟨f, hf⟩ have F : ∀ n, f n ∈ S := fun n => by rw [hf]; exact mem_range_self _ rw [hf, sInter_range] - exact dense_iInter_of_open_nat (fun n => ho _ (F n)) fun n => hd _ (F n) -#align dense_sInter_of_open dense_sInter_of_open + exact dense_iInter_of_isOpen_nat (fun n => ho _ (F n)) fun n => hd _ (F n) +#align dense_sInter_of_open dense_sInter_of_isOpen /-- Baire theorem: a countable intersection of dense open sets is dense. Formulated here with an index set which is a countable set in any type. -/ -theorem dense_biInter_of_open {S : Set β} {f : β → Set α} (ho : ∀ s ∈ S, IsOpen (f s)) +theorem dense_biInter_of_isOpen {S : Set β} {f : β → Set α} (ho : ∀ s ∈ S, IsOpen (f s)) (hS : S.Countable) (hd : ∀ s ∈ S, Dense (f s)) : Dense (⋂ s ∈ S, f s) := by rw [← sInter_image] - apply dense_sInter_of_open + apply dense_sInter_of_isOpen · rwa [ball_image_iff] · exact hS.image _ · rwa [ball_image_iff] -#align dense_bInter_of_open dense_biInter_of_open +#align dense_bInter_of_open dense_biInter_of_isOpen /-- Baire theorem: a countable intersection of dense open sets is dense. Formulated here with an index set which is an encodable type. -/ -theorem dense_iInter_of_open [Encodable β] {f : β → Set α} (ho : ∀ s, IsOpen (f s)) +theorem dense_iInter_of_isOpen [Encodable β] {f : β → Set α} (ho : ∀ s, IsOpen (f s)) (hd : ∀ s, Dense (f s)) : Dense (⋂ s, f s) := by rw [← sInter_range] - apply dense_sInter_of_open + apply dense_sInter_of_isOpen · rwa [forall_range_iff] · exact countable_range _ · rwa [forall_range_iff] -#align dense_Inter_of_open dense_iInter_of_open +#align dense_Inter_of_open dense_iInter_of_isOpen /-- A set is residual (comeagre) if and only if it includes a dense `Gδ` set. -/ theorem mem_residual {s : Set α} : s ∈ residual α ↔ ∃ (t : _) (_ : t ⊆ s), IsGδ t ∧ Dense t := by @@ -231,7 +231,7 @@ theorem mem_residual {s : Set α} : s ∈ residual α ↔ ∃ (t : _) (_ : t ⊆ · rw [mem_residual_iff] rintro ⟨S, hSo, hSd, Sct, Ss⟩ refine' ⟨_, Ss, ⟨_, fun t ht => hSo _ ht, Sct, rfl⟩, _⟩ - exact dense_sInter_of_open hSo Sct hSd + exact dense_sInter_of_isOpen hSo Sct hSd rintro ⟨t, ts, ho, hd⟩ exact mem_of_superset (residual_of_dense_Gδ ho hd) ts #align mem_residual mem_residual @@ -294,10 +294,10 @@ theorem IsGδ.dense_iUnion_interior_of_closed [Encodable ι] {s : Set α} (hs : let g i := (frontier (f i))ᶜ have hgo : ∀ i, IsOpen (g i) := fun i => isClosed_frontier.isOpen_compl have hgd : Dense (⋂ i, g i) := by - refine' dense_iInter_of_open hgo fun i x => _ + refine' dense_iInter_of_isOpen hgo fun i x => _ rw [closure_compl, interior_frontier (hc _)] exact id - refine' (hd.inter_of_Gδ hs (isGδ_iInter_of_open fun i => (hgo i)) hgd).mono _ + refine' (hd.inter_of_Gδ hs (isGδ_iInter_of_isOpen fun i => (hgo i)) hgd).mono _ rintro x ⟨hxs, hxg⟩ rw [mem_iInter] at hxg rcases mem_iUnion.1 (hU hxs) with ⟨i, hi⟩ diff --git a/Mathlib/Topology/MetricSpace/DilationEquiv.lean b/Mathlib/Topology/MetricSpace/DilationEquiv.lean index 181a3879eec8e..b2401d76d0f76 100644 --- a/Mathlib/Topology/MetricSpace/DilationEquiv.lean +++ b/Mathlib/Topology/MetricSpace/DilationEquiv.lean @@ -77,6 +77,10 @@ def symm (e : X ≃ᵈ Y) : Y ≃ᵈ X where ENNReal.coe_one, one_mul] @[simp] theorem symm_symm (e : X ≃ᵈ Y) : e.symm.symm = e := rfl + +theorem symm_bijective : Function.Bijective (DilationEquiv.symm : (X ≃ᵈ Y) → Y ≃ᵈ X) := + Function.bijective_iff_has_inverse.mpr ⟨_, symm_symm, symm_symm⟩ + @[simp] theorem apply_symm_apply (e : X ≃ᵈ Y) (x : Y) : e (e.symm x) = x := e.right_inv x @[simp] theorem symm_apply_apply (e : X ≃ᵈ Y) (x : X) : e.symm (e x) = x := e.left_inv x diff --git a/Mathlib/Topology/MetricSpace/Isometry.lean b/Mathlib/Topology/MetricSpace/Isometry.lean index aad5693c4a957..bf96d3fb0d64b 100644 --- a/Mathlib/Topology/MetricSpace/Isometry.lean +++ b/Mathlib/Topology/MetricSpace/Isometry.lean @@ -420,6 +420,9 @@ initialize_simps_projections IsometryEquiv (toEquiv_toFun → apply, toEquiv_inv theorem symm_symm (h : α ≃ᵢ β) : h.symm.symm = h := rfl #align isometry_equiv.symm_symm IsometryEquiv.symm_symm +theorem symm_bijective : Bijective (IsometryEquiv.symm : (α ≃ᵢ β) → β ≃ᵢ α) := + Function.bijective_iff_has_inverse.mpr ⟨_, symm_symm, symm_symm⟩ + @[simp] theorem apply_symm_apply (h : α ≃ᵢ β) (y : β) : h (h.symm y) = y := h.toEquiv.apply_symm_apply y diff --git a/Mathlib/Topology/MetricSpace/Kuratowski.lean b/Mathlib/Topology/MetricSpace/Kuratowski.lean index 18695c7976d93..17bd6111f318e 100644 --- a/Mathlib/Topology/MetricSpace/Kuratowski.lean +++ b/Mathlib/Topology/MetricSpace/Kuratowski.lean @@ -155,7 +155,7 @@ theorem LipschitzOnWith.extend_lp_infty [PseudoMetricSpace α] {s : Set α} {f : · apply LipschitzWith.uniformly_bounded (swap g) hgl a₀ use ‖f a₀‖ rintro - ⟨i, rfl⟩ - simp_rw [←hgeq i ha₀_in_s] + simp_rw [← hgeq i ha₀_in_s] exact lp.norm_apply_le_norm top_ne_zero (f a₀) i -- Construct witness by bundling the function with its certificate of membership in ℓ^∞ let f_ext' : α → ℓ^∞(ι) := fun i ↦ ⟨swap g i, hf_extb i⟩ diff --git a/Mathlib/Topology/MetricSpace/PiNat.lean b/Mathlib/Topology/MetricSpace/PiNat.lean index 73f9f6d665dd2..88af496faa80c 100644 --- a/Mathlib/Topology/MetricSpace/PiNat.lean +++ b/Mathlib/Topology/MetricSpace/PiNat.lean @@ -90,7 +90,7 @@ theorem firstDiff_comm (x y : ∀ n, E n) : firstDiff x y = firstDiff y x := by theorem min_firstDiff_le (x y z : ∀ n, E n) (h : x ≠ z) : min (firstDiff x y) (firstDiff y z) ≤ firstDiff x z := by - by_contra' H + by_contra! H rw [lt_min_iff] at H refine apply_firstDiff_ne h ?_ calc @@ -154,7 +154,7 @@ theorem mem_cylinder_iff_le_firstDiff {x y : ∀ n, E n} (hne : x ≠ y) (i : x ∈ cylinder y i ↔ i ≤ firstDiff x y := by constructor · intro h - by_contra' + by_contra! exact apply_firstDiff_ne hne (h _ this) · intro hi j hj exact apply_eq_of_lt_firstDiff (hj.trans_le hi) @@ -317,7 +317,7 @@ theorem mem_cylinder_iff_dist_le {x y : ∀ n, E n} {n : ℕ} : suffices (∀ i : ℕ, i < n → y i = x i) ↔ n ≤ firstDiff y x by simpa [dist_eq_of_ne hne] constructor · intro hy - by_contra' H + by_contra! H exact apply_firstDiff_ne hne (hy _ H) · intro h i hi exact apply_eq_of_lt_firstDiff (hi.trans_le h) @@ -363,7 +363,7 @@ theorem isOpen_cylinder (x : ∀ n, E n) (n : ℕ) : IsOpen (cylinder x n) := by theorem isTopologicalBasis_cylinders : IsTopologicalBasis { s : Set (∀ n, E n) | ∃ (x : ∀ n, E n) (n : ℕ), s = cylinder x n } := by - apply isTopologicalBasis_of_open_of_nhds + apply isTopologicalBasis_of_isOpen_of_nhds · rintro u ⟨x, n, rfl⟩ apply isOpen_cylinder · intro x u hx u_open @@ -744,7 +744,7 @@ theorem exists_nat_nat_continuous_surjective_of_completeSpace (α : Type*) [Metr have dist' : dist x y = dist x.1 y.1 := rfl let n := firstDiff x.1 y.1 - 1 have diff_pos : 0 < firstDiff x.1 y.1 := by - by_contra' h + by_contra! h apply apply_firstDiff_ne hne' rw [le_zero_iff.1 h] apply apply_eq_of_dist_lt _ le_rfl diff --git a/Mathlib/Topology/MetricSpace/ThickenedIndicator.lean b/Mathlib/Topology/MetricSpace/ThickenedIndicator.lean index 53aadaba0b7f8..49193fec78031 100644 --- a/Mathlib/Topology/MetricSpace/ThickenedIndicator.lean +++ b/Mathlib/Topology/MetricSpace/ThickenedIndicator.lean @@ -196,6 +196,16 @@ theorem thickenedIndicator_one_of_mem_closure {δ : ℝ} (δ_pos : 0 < δ) (E : rw [thickenedIndicator_apply, thickenedIndicatorAux_one_of_mem_closure δ E x_mem, one_toNNReal] #align thickened_indicator_one_of_mem_closure thickenedIndicator_one_of_mem_closure +lemma one_le_thickenedIndicator_apply' {X : Type _} [PseudoEMetricSpace X] + {δ : ℝ} (δ_pos : 0 < δ) {F : Set X} {x : X} (hxF : x ∈ closure F) : + 1 ≤ thickenedIndicator δ_pos F x := by + rw [thickenedIndicator_one_of_mem_closure δ_pos F hxF] + +lemma one_le_thickenedIndicator_apply (X : Type _) [PseudoEMetricSpace X] + {δ : ℝ} (δ_pos : 0 < δ) {F : Set X} {x : X} (hxF : x ∈ F) : + 1 ≤ thickenedIndicator δ_pos F x := + one_le_thickenedIndicator_apply' δ_pos (subset_closure hxF) + theorem thickenedIndicator_one {δ : ℝ} (δ_pos : 0 < δ) (E : Set α) {x : α} (x_in_E : x ∈ E) : thickenedIndicator δ_pos E x = 1 := thickenedIndicator_one_of_mem_closure _ _ (subset_closure x_in_E) diff --git a/Mathlib/Topology/NhdsSet.lean b/Mathlib/Topology/NhdsSet.lean index 82ae64bbfa696..919f394686387 100644 --- a/Mathlib/Topology/NhdsSet.lean +++ b/Mathlib/Topology/NhdsSet.lean @@ -175,5 +175,5 @@ theorem Continuous.tendsto_nhdsSet {f : α → β} {t : Set β} (hf : Continuous lemma Continuous.tendsto_nhdsSet_nhds {X Y : Type*} [TopologicalSpace X] [TopologicalSpace Y] {s : Set X} {y : Y} {f : X → Y} (h : Continuous f) (h' : EqOn f (fun _ ↦ y) s) : Tendsto f (𝓝ˢ s) (𝓝 y) := by - rw [←nhdsSet_singleton] + rw [← nhdsSet_singleton] exact h.tendsto_nhdsSet h' diff --git a/Mathlib/Topology/Order/Basic.lean b/Mathlib/Topology/Order/Basic.lean index 4df33c7c2e5f3..1e1b84e6a24ad 100644 --- a/Mathlib/Topology/Order/Basic.lean +++ b/Mathlib/Topology/Order/Basic.lean @@ -1365,11 +1365,11 @@ theorem countable_setOf_covby_right [SecondCountableTopology α] : rcases hxx'.lt_or_lt with (h' | h') · refine' disjoint_left.2 fun u ux ux' => xt.2.2.1 _ refine' h'z x' x't ⟨ux'.1.trans_le (ux.2.trans (hy x xt.1).le), _⟩ - by_contra' H + by_contra! H exact lt_irrefl _ ((Hy _ _ xt.1 H).trans_lt h') · refine' disjoint_left.2 fun u ux ux' => x't.2.2.1 _ refine' h'z x xt ⟨ux.1.trans_le (ux'.2.trans (hy x' x't.1).le), _⟩ - by_contra' H + by_contra! H exact lt_irrefl _ ((Hy _ _ x't.1 H).trans_lt h') refine' this.countable_of_isOpen (fun x hx => _) fun x hx => ⟨x, hz x hx, le_rfl⟩ suffices H : Ioc (z x) x = Ioo (z x) (y x) diff --git a/Mathlib/Topology/Order/Lattice.lean b/Mathlib/Topology/Order/Lattice.lean index 7713724aa8206..d83c889d90701 100644 --- a/Mathlib/Topology/Order/Lattice.lean +++ b/Mathlib/Topology/Order/Lattice.lean @@ -26,7 +26,6 @@ topological, lattice set_option autoImplicit true - open Filter open Topology @@ -75,8 +74,7 @@ instance (priority := 100) OrderDual.topologicalLattice (L : Type*) [Topological -- see Note [lower instance priority] instance (priority := 100) LinearOrder.topologicalLattice {L : Type*} [TopologicalSpace L] - [LinearOrder L] [OrderClosedTopology L] : TopologicalLattice L - where + [LinearOrder L] [OrderClosedTopology L] : TopologicalLattice L where continuous_inf := continuous_min continuous_sup := continuous_max #align linear_order.topological_lattice LinearOrder.topologicalLattice @@ -105,26 +103,308 @@ theorem Continuous.sup [Sup L] [ContinuousSup L] {f g : X → L} (hf : Continuou continuous_sup.comp (hf.prod_mk hg : _) #align continuous.sup Continuous.sup -theorem Filter.Tendsto.sup_right_nhds' {ι β} [TopologicalSpace β] [Sup β] [ContinuousSup β] - {l : Filter ι} {f g : ι → β} {x y : β} (hf : Tendsto f l (𝓝 x)) (hg : Tendsto g l (𝓝 y)) : +namespace Filter.Tendsto + +section SupInf + +variable {α : Type*} {l : Filter α} {f g : α → L} {x y : L} + +lemma sup_nhds' [Sup L] [ContinuousSup L] (hf : Tendsto f l (𝓝 x)) (hg : Tendsto g l (𝓝 y)) : Tendsto (f ⊔ g) l (𝓝 (x ⊔ y)) := (continuous_sup.tendsto _).comp (Tendsto.prod_mk_nhds hf hg) -#align filter.tendsto.sup_right_nhds' Filter.Tendsto.sup_right_nhds' +#align filter.tendsto.sup_right_nhds' Filter.Tendsto.sup_nhds' -theorem Filter.Tendsto.sup_right_nhds {ι β} [TopologicalSpace β] [Sup β] [ContinuousSup β] - {l : Filter ι} {f g : ι → β} {x y : β} (hf : Tendsto f l (𝓝 x)) (hg : Tendsto g l (𝓝 y)) : +lemma sup_nhds [Sup L] [ContinuousSup L] (hf : Tendsto f l (𝓝 x)) (hg : Tendsto g l (𝓝 y)) : Tendsto (fun i => f i ⊔ g i) l (𝓝 (x ⊔ y)) := - hf.sup_right_nhds' hg -#align filter.tendsto.sup_right_nhds Filter.Tendsto.sup_right_nhds + hf.sup_nhds' hg +#align filter.tendsto.sup_right_nhds Filter.Tendsto.sup_nhds -theorem Filter.Tendsto.inf_right_nhds' {ι β} [TopologicalSpace β] [Inf β] [ContinuousInf β] - {l : Filter ι} {f g : ι → β} {x y : β} (hf : Tendsto f l (𝓝 x)) (hg : Tendsto g l (𝓝 y)) : +lemma inf_nhds' [Inf L] [ContinuousInf L] (hf : Tendsto f l (𝓝 x)) (hg : Tendsto g l (𝓝 y)) : Tendsto (f ⊓ g) l (𝓝 (x ⊓ y)) := (continuous_inf.tendsto _).comp (Tendsto.prod_mk_nhds hf hg) -#align filter.tendsto.inf_right_nhds' Filter.Tendsto.inf_right_nhds' +#align filter.tendsto.inf_right_nhds' Filter.Tendsto.inf_nhds' -theorem Filter.Tendsto.inf_right_nhds {ι β} [TopologicalSpace β] [Inf β] [ContinuousInf β] - {l : Filter ι} {f g : ι → β} {x y : β} (hf : Tendsto f l (𝓝 x)) (hg : Tendsto g l (𝓝 y)) : +lemma inf_nhds [Inf L] [ContinuousInf L] (hf : Tendsto f l (𝓝 x)) (hg : Tendsto g l (𝓝 y)) : Tendsto (fun i => f i ⊓ g i) l (𝓝 (x ⊓ y)) := - hf.inf_right_nhds' hg -#align filter.tendsto.inf_right_nhds Filter.Tendsto.inf_right_nhds + hf.inf_nhds' hg +#align filter.tendsto.inf_right_nhds Filter.Tendsto.inf_nhds + +end SupInf + +open Finset + +variable {ι : Type*} {s : Finset ι} {f : ι → α → L} {g : ι → L} + +lemma finset_sup'_nhds [SemilatticeSup L] [ContinuousSup L] + (hne : s.Nonempty) (hs : ∀ i ∈ s, Tendsto (f i) l (𝓝 (g i))) : + Tendsto (s.sup' hne f) l (𝓝 (s.sup' hne g)) := by + induction hne using Finset.Nonempty.cons_induction with + | h₀ => simpa using hs + | h₁ s ha hne ihs => + rw [forall_mem_cons] at hs + simp only [sup'_cons, hne] + exact hs.1.sup_nhds (ihs hs.2) + +lemma finset_sup'_nhds_apply [SemilatticeSup L] [ContinuousSup L] + (hne : s.Nonempty) (hs : ∀ i ∈ s, Tendsto (f i) l (𝓝 (g i))) : + Tendsto (fun a ↦ s.sup' hne (f · a)) l (𝓝 (s.sup' hne g)) := by + simpa only [← Finset.sup'_apply] using finset_sup'_nhds hne hs + +lemma finset_inf'_nhds [SemilatticeInf L] [ContinuousInf L] + (hne : s.Nonempty) (hs : ∀ i ∈ s, Tendsto (f i) l (𝓝 (g i))) : + Tendsto (s.inf' hne f) l (𝓝 (s.inf' hne g)) := + finset_sup'_nhds (L := Lᵒᵈ) hne hs + +lemma finset_inf'_nhds_apply [SemilatticeInf L] [ContinuousInf L] + (hne : s.Nonempty) (hs : ∀ i ∈ s, Tendsto (f i) l (𝓝 (g i))) : + Tendsto (fun a ↦ s.inf' hne (f · a)) l (𝓝 (s.inf' hne g)) := + finset_sup'_nhds_apply (L := Lᵒᵈ) hne hs + +lemma finset_sup_nhds [SemilatticeSup L] [OrderBot L] [ContinuousSup L] + (hs : ∀ i ∈ s, Tendsto (f i) l (𝓝 (g i))) : Tendsto (s.sup f) l (𝓝 (s.sup g)) := by + rcases s.eq_empty_or_nonempty with rfl | hne + · simpa using tendsto_const_nhds + · simp only [← sup'_eq_sup hne] + exact finset_sup'_nhds hne hs + +lemma finset_sup_nhds_apply [SemilatticeSup L] [OrderBot L] [ContinuousSup L] + (hs : ∀ i ∈ s, Tendsto (f i) l (𝓝 (g i))) : + Tendsto (fun a ↦ s.sup (f · a)) l (𝓝 (s.sup g)) := by + simpa only [← Finset.sup_apply] using finset_sup_nhds hs + +lemma finset_inf_nhds [SemilatticeInf L] [OrderTop L] [ContinuousInf L] + (hs : ∀ i ∈ s, Tendsto (f i) l (𝓝 (g i))) : Tendsto (s.inf f) l (𝓝 (s.inf g)) := + finset_sup_nhds (L := Lᵒᵈ) hs + +lemma finset_inf_nhds_apply [SemilatticeInf L] [OrderTop L] [ContinuousInf L] + (hs : ∀ i ∈ s, Tendsto (f i) l (𝓝 (g i))) : + Tendsto (fun a ↦ s.inf (f · a)) l (𝓝 (s.inf g)) := + finset_sup_nhds_apply (L := Lᵒᵈ) hs + +end Filter.Tendsto + +section Sup + +variable [Sup L] [ContinuousSup L] {f g : X → L} {s : Set X} {x : X} + +lemma ContinuousAt.sup' (hf : ContinuousAt f x) (hg : ContinuousAt g x) : + ContinuousAt (f ⊔ g) x := + hf.sup_nhds' hg + +lemma ContinuousAt.sup (hf : ContinuousAt f x) (hg : ContinuousAt g x) : + ContinuousAt (fun a ↦ f a ⊔ g a) x := + hf.sup' hg + +lemma ContinuousWithinAt.sup' (hf : ContinuousWithinAt f s x) (hg : ContinuousWithinAt g s x) : + ContinuousWithinAt (f ⊔ g) s x := + hf.sup_nhds' hg + +lemma ContinuousWithinAt.sup (hf : ContinuousWithinAt f s x) (hg : ContinuousWithinAt g s x) : + ContinuousWithinAt (fun a ↦ f a ⊔ g a) s x := + hf.sup' hg + +lemma ContinuousOn.sup' (hf : ContinuousOn f s) (hg : ContinuousOn g s) : + ContinuousOn (f ⊔ g) s := fun x hx ↦ + (hf x hx).sup' (hg x hx) + +lemma ContinuousOn.sup (hf : ContinuousOn f s) (hg : ContinuousOn g s) : + ContinuousOn (fun a ↦ f a ⊔ g a) s := + hf.sup' hg + +lemma Continuous.sup' (hf : Continuous f) (hg : Continuous g) : Continuous (f ⊔ g) := hf.sup hg + +end Sup + +section Inf + +variable [Inf L] [ContinuousInf L] {f g : X → L} {s : Set X} {x : X} + +lemma ContinuousAt.inf' (hf : ContinuousAt f x) (hg : ContinuousAt g x) : + ContinuousAt (f ⊓ g) x := + hf.inf_nhds' hg + +lemma ContinuousAt.inf (hf : ContinuousAt f x) (hg : ContinuousAt g x) : + ContinuousAt (fun a ↦ f a ⊓ g a) x := + hf.inf' hg + +lemma ContinuousWithinAt.inf' (hf : ContinuousWithinAt f s x) (hg : ContinuousWithinAt g s x) : + ContinuousWithinAt (f ⊓ g) s x := + hf.inf_nhds' hg + +lemma ContinuousWithinAt.inf (hf : ContinuousWithinAt f s x) (hg : ContinuousWithinAt g s x) : + ContinuousWithinAt (fun a ↦ f a ⊓ g a) s x := + hf.inf' hg + +lemma ContinuousOn.inf' (hf : ContinuousOn f s) (hg : ContinuousOn g s) : + ContinuousOn (f ⊓ g) s := fun x hx ↦ + (hf x hx).inf' (hg x hx) + +lemma ContinuousOn.inf (hf : ContinuousOn f s) (hg : ContinuousOn g s) : + ContinuousOn (fun a ↦ f a ⊓ g a) s := + hf.inf' hg + +lemma Continuous.inf' (hf : Continuous f) (hg : Continuous g) : Continuous (f ⊓ g) := hf.inf hg + +end Inf + +section FinsetSup' + +variable {ι : Type*} [SemilatticeSup L] [ContinuousSup L] {s : Finset ι} + {f : ι → X → L} {t : Set X} {x : X} + +lemma ContinuousAt.finset_sup'_apply (hne : s.Nonempty) (hs : ∀ i ∈ s, ContinuousAt (f i) x) : + ContinuousAt (fun a ↦ s.sup' hne (f · a)) x := + Tendsto.finset_sup'_nhds_apply hne hs + +lemma ContinuousAt.finset_sup' (hne : s.Nonempty) (hs : ∀ i ∈ s, ContinuousAt (f i) x) : + ContinuousAt (s.sup' hne f) x := by + simpa only [← Finset.sup'_apply] using finset_sup'_apply hne hs + +lemma ContinuousWithinAt.finset_sup'_apply (hne : s.Nonempty) + (hs : ∀ i ∈ s, ContinuousWithinAt (f i) t x) : + ContinuousWithinAt (fun a ↦ s.sup' hne (f · a)) t x := + Tendsto.finset_sup'_nhds_apply hne hs + +lemma ContinuousWithinAt.finset_sup' (hne : s.Nonempty) + (hs : ∀ i ∈ s, ContinuousWithinAt (f i) t x) : ContinuousWithinAt (s.sup' hne f) t x := by + simpa only [← Finset.sup'_apply] using finset_sup'_apply hne hs + +lemma ContinuousOn.finset_sup'_apply (hne : s.Nonempty) (hs : ∀ i ∈ s, ContinuousOn (f i) t) : + ContinuousOn (fun a ↦ s.sup' hne (f · a)) t := fun x hx ↦ + ContinuousWithinAt.finset_sup'_apply hne fun i hi ↦ hs i hi x hx + +lemma ContinuousOn.finset_sup' (hne : s.Nonempty) (hs : ∀ i ∈ s, ContinuousOn (f i) t) : + ContinuousOn (s.sup' hne f) t := fun x hx ↦ + ContinuousWithinAt.finset_sup' hne fun i hi ↦ hs i hi x hx + +lemma Continuous.finset_sup'_apply (hne : s.Nonempty) (hs : ∀ i ∈ s, Continuous (f i)) : + Continuous (fun a ↦ s.sup' hne (f · a)) := + continuous_iff_continuousAt.2 fun _ ↦ ContinuousAt.finset_sup'_apply _ fun i hi ↦ + (hs i hi).continuousAt + +lemma Continuous.finset_sup' (hne : s.Nonempty) (hs : ∀ i ∈ s, Continuous (f i)) : + Continuous (s.sup' hne f) := + continuous_iff_continuousAt.2 fun _ ↦ ContinuousAt.finset_sup' _ fun i hi ↦ (hs i hi).continuousAt + +end FinsetSup' + +section FinsetSup + +variable {ι : Type*} [SemilatticeSup L] [OrderBot L] [ContinuousSup L] {s : Finset ι} + {f : ι → X → L} {t : Set X} {x : X} + +lemma ContinuousAt.finset_sup_apply (hs : ∀ i ∈ s, ContinuousAt (f i) x) : + ContinuousAt (fun a ↦ s.sup (f · a)) x := + Tendsto.finset_sup_nhds_apply hs + +lemma ContinuousAt.finset_sup (hs : ∀ i ∈ s, ContinuousAt (f i) x) : + ContinuousAt (s.sup f) x := by + simpa only [← Finset.sup_apply] using finset_sup_apply hs + +lemma ContinuousWithinAt.finset_sup_apply + (hs : ∀ i ∈ s, ContinuousWithinAt (f i) t x) : + ContinuousWithinAt (fun a ↦ s.sup (f · a)) t x := + Tendsto.finset_sup_nhds_apply hs + +lemma ContinuousWithinAt.finset_sup + (hs : ∀ i ∈ s, ContinuousWithinAt (f i) t x) : ContinuousWithinAt (s.sup f) t x := by + simpa only [← Finset.sup_apply] using finset_sup_apply hs + +lemma ContinuousOn.finset_sup_apply (hs : ∀ i ∈ s, ContinuousOn (f i) t) : + ContinuousOn (fun a ↦ s.sup (f · a)) t := fun x hx ↦ + ContinuousWithinAt.finset_sup_apply fun i hi ↦ hs i hi x hx + +lemma ContinuousOn.finset_sup (hs : ∀ i ∈ s, ContinuousOn (f i) t) : + ContinuousOn (s.sup f) t := fun x hx ↦ + ContinuousWithinAt.finset_sup fun i hi ↦ hs i hi x hx + +lemma Continuous.finset_sup_apply (hs : ∀ i ∈ s, Continuous (f i)) : + Continuous (fun a ↦ s.sup (f · a)) := + continuous_iff_continuousAt.2 fun _ ↦ ContinuousAt.finset_sup_apply fun i hi ↦ + (hs i hi).continuousAt + +lemma Continuous.finset_sup (hs : ∀ i ∈ s, Continuous (f i)) : Continuous (s.sup f) := + continuous_iff_continuousAt.2 fun _ ↦ ContinuousAt.finset_sup fun i hi ↦ (hs i hi).continuousAt + +end FinsetSup + +section FinsetInf' + +variable {ι : Type*} [SemilatticeInf L] [ContinuousInf L] {s : Finset ι} + {f : ι → X → L} {t : Set X} {x : X} + +lemma ContinuousAt.finset_inf'_apply (hne : s.Nonempty) (hs : ∀ i ∈ s, ContinuousAt (f i) x) : + ContinuousAt (fun a ↦ s.inf' hne (f · a)) x := + Tendsto.finset_inf'_nhds_apply hne hs + +lemma ContinuousAt.finset_inf' (hne : s.Nonempty) (hs : ∀ i ∈ s, ContinuousAt (f i) x) : + ContinuousAt (s.inf' hne f) x := by + simpa only [← Finset.inf'_apply] using finset_inf'_apply hne hs + +lemma ContinuousWithinAt.finset_inf'_apply (hne : s.Nonempty) + (hs : ∀ i ∈ s, ContinuousWithinAt (f i) t x) : + ContinuousWithinAt (fun a ↦ s.inf' hne (f · a)) t x := + Tendsto.finset_inf'_nhds_apply hne hs + +lemma ContinuousWithinAt.finset_inf' (hne : s.Nonempty) + (hs : ∀ i ∈ s, ContinuousWithinAt (f i) t x) : ContinuousWithinAt (s.inf' hne f) t x := by + simpa only [← Finset.inf'_apply] using finset_inf'_apply hne hs + +lemma ContinuousOn.finset_inf'_apply (hne : s.Nonempty) (hs : ∀ i ∈ s, ContinuousOn (f i) t) : + ContinuousOn (fun a ↦ s.inf' hne (f · a)) t := fun x hx ↦ + ContinuousWithinAt.finset_inf'_apply hne fun i hi ↦ hs i hi x hx + +lemma ContinuousOn.finset_inf' (hne : s.Nonempty) (hs : ∀ i ∈ s, ContinuousOn (f i) t) : + ContinuousOn (s.inf' hne f) t := fun x hx ↦ + ContinuousWithinAt.finset_inf' hne fun i hi ↦ hs i hi x hx + +lemma Continuous.finset_inf'_apply (hne : s.Nonempty) (hs : ∀ i ∈ s, Continuous (f i)) : + Continuous (fun a ↦ s.inf' hne (f · a)) := + continuous_iff_continuousAt.2 fun _ ↦ ContinuousAt.finset_inf'_apply _ fun i hi ↦ + (hs i hi).continuousAt + +lemma Continuous.finset_inf' (hne : s.Nonempty) (hs : ∀ i ∈ s, Continuous (f i)) : + Continuous (s.inf' hne f) := + continuous_iff_continuousAt.2 fun _ ↦ ContinuousAt.finset_inf' _ fun i hi ↦ (hs i hi).continuousAt + +end FinsetInf' + +section FinsetInf + +variable {ι : Type*} [SemilatticeInf L] [OrderTop L] [ContinuousInf L] {s : Finset ι} + {f : ι → X → L} {t : Set X} {x : X} + +lemma ContinuousAt.finset_inf_apply (hs : ∀ i ∈ s, ContinuousAt (f i) x) : + ContinuousAt (fun a ↦ s.inf (f · a)) x := + Tendsto.finset_inf_nhds_apply hs + +lemma ContinuousAt.finset_inf (hs : ∀ i ∈ s, ContinuousAt (f i) x) : + ContinuousAt (s.inf f) x := by + simpa only [← Finset.inf_apply] using finset_inf_apply hs + +lemma ContinuousWithinAt.finset_inf_apply + (hs : ∀ i ∈ s, ContinuousWithinAt (f i) t x) : + ContinuousWithinAt (fun a ↦ s.inf (f · a)) t x := + Tendsto.finset_inf_nhds_apply hs + +lemma ContinuousWithinAt.finset_inf + (hs : ∀ i ∈ s, ContinuousWithinAt (f i) t x) : ContinuousWithinAt (s.inf f) t x := by + simpa only [← Finset.inf_apply] using finset_inf_apply hs + +lemma ContinuousOn.finset_inf_apply (hs : ∀ i ∈ s, ContinuousOn (f i) t) : + ContinuousOn (fun a ↦ s.inf (f · a)) t := fun x hx ↦ + ContinuousWithinAt.finset_inf_apply fun i hi ↦ hs i hi x hx + +lemma ContinuousOn.finset_inf (hs : ∀ i ∈ s, ContinuousOn (f i) t) : + ContinuousOn (s.inf f) t := fun x hx ↦ + ContinuousWithinAt.finset_inf fun i hi ↦ hs i hi x hx + +lemma Continuous.finset_inf_apply (hs : ∀ i ∈ s, Continuous (f i)) : + Continuous (fun a ↦ s.inf (f · a)) := + continuous_iff_continuousAt.2 fun _ ↦ ContinuousAt.finset_inf_apply fun i hi ↦ + (hs i hi).continuousAt + +lemma Continuous.finset_inf (hs : ∀ i ∈ s, Continuous (f i)) : Continuous (s.inf f) := + continuous_iff_continuousAt.2 fun _ ↦ ContinuousAt.finset_inf fun i hi ↦ (hs i hi).continuousAt + +end FinsetInf diff --git a/Mathlib/Topology/Order/PartialSups.lean b/Mathlib/Topology/Order/PartialSups.lean new file mode 100644 index 0000000000000..954b72b9b03aa --- /dev/null +++ b/Mathlib/Topology/Order/PartialSups.lean @@ -0,0 +1,70 @@ +/- +Copyright (c) 2023 Yury Kudryashov. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yury Kudryashov +-/ +import Mathlib.Topology.Order.Lattice +import Mathlib.Order.PartialSups + +/-! +# Continuity of `partialSups` + +In this file we prove that `partialSups` of a sequence of continuous functions is continuous +as well as versions for `Filter.Tendsto`, `ContinuousAt`, `ContinuousWithinAt`, and `ContinuousOn`. +-/ + +open Filter +open scoped Topology + +variable {L : Type*} [SemilatticeSup L] [TopologicalSpace L] [ContinuousSup L] + +namespace Filter.Tendsto + +variable {α : Type*} {l : Filter α} {f : ℕ → α → L} {g : ℕ → L} {n : ℕ} + +protected lemma partialSups (hf : ∀ k ≤ n, Tendsto (f k) l (𝓝 (g k))) : + Tendsto (partialSups f n) l (𝓝 (partialSups g n)) := by + simp only [partialSups_eq_sup'_range] + refine finset_sup'_nhds _ ?_ + simpa [Nat.lt_succ_iff] + +protected lemma partialSups_apply (hf : ∀ k ≤ n, Tendsto (f k) l (𝓝 (g k))) : + Tendsto (fun a ↦ partialSups (f · a) n) l (𝓝 (partialSups g n)) := by + simpa only [← partialSups_apply] using Tendsto.partialSups hf + +end Filter.Tendsto + +variable {X : Type*} [TopologicalSpace X] {f : ℕ → X → L} {n : ℕ} {s : Set X} {x : X} + +protected lemma ContinuousAt.partialSups_apply (hf : ∀ k ≤ n, ContinuousAt (f k) x) : + ContinuousAt (fun a ↦ partialSups (f · a) n) x := + Tendsto.partialSups_apply hf + +protected lemma ContinuousAt.partialSups (hf : ∀ k ≤ n, ContinuousAt (f k) x) : + ContinuousAt (partialSups f n) x := by + simpa only [← partialSups_apply] using ContinuousAt.partialSups_apply hf + +protected lemma ContinuousWithinAt.partialSups_apply (hf : ∀ k ≤ n, ContinuousWithinAt (f k) s x) : + ContinuousWithinAt (fun a ↦ partialSups (f · a) n) s x := + Tendsto.partialSups_apply hf + +protected lemma ContinuousWithinAt.partialSups (hf : ∀ k ≤ n, ContinuousWithinAt (f k) s x) : + ContinuousWithinAt (partialSups f n) s x := by + simpa only [← partialSups_apply] using ContinuousWithinAt.partialSups_apply hf + +protected lemma ContinuousOn.partialSups_apply (hf : ∀ k ≤ n, ContinuousOn (f k) s) : + ContinuousOn (fun a ↦ partialSups (f · a) n) s := fun x hx ↦ + ContinuousWithinAt.partialSups_apply fun k hk ↦ hf k hk x hx + +protected lemma ContinuousOn.partialSups (hf : ∀ k ≤ n, ContinuousOn (f k) s) : + ContinuousOn (partialSups f n) s := fun x hx ↦ + ContinuousWithinAt.partialSups fun k hk ↦ hf k hk x hx + +protected lemma Continuous.partialSups_apply (hf : ∀ k ≤ n, Continuous (f k)) : + Continuous (fun a ↦ partialSups (f · a) n) := + continuous_iff_continuousAt.2 fun _ ↦ ContinuousAt.partialSups_apply fun k hk ↦ + (hf k hk).continuousAt + +protected lemma Continuous.partialSups (hf : ∀ k ≤ n, Continuous (f k)) : + Continuous (partialSups f n) := + continuous_iff_continuousAt.2 fun _ ↦ ContinuousAt.partialSups fun k hk ↦ (hf k hk).continuousAt diff --git a/Mathlib/Topology/Perfect.lean b/Mathlib/Topology/Perfect.lean index 5f7ee247574a9..a9b0a03418a19 100644 --- a/Mathlib/Topology/Perfect.lean +++ b/Mathlib/Topology/Perfect.lean @@ -191,7 +191,7 @@ theorem exists_countable_union_perfect_of_isClosed [SecondCountableTopology α] have : U ∈ v := ⟨hUb, hU_cnt⟩ apply xD.2 exact mem_biUnion this xU - by_contra' h + by_contra! h exact absurd (Countable.mono h (Set.countable_singleton _)) this · rw [inter_comm, inter_union_diff] #align exists_countable_union_perfect_of_is_closed exists_countable_union_perfect_of_isClosed diff --git a/Mathlib/Topology/Separation.lean b/Mathlib/Topology/Separation.lean index a1644023433d4..bccc397e77bd7 100644 --- a/Mathlib/Topology/Separation.lean +++ b/Mathlib/Topology/Separation.lean @@ -44,9 +44,9 @@ This file defines the predicate `SeparatedNhds`, and common separation axioms ### T₀ spaces -* `IsClosed.exists_closed_singleton` Given a closed set `S` in a compact T₀ space, +* `IsClosed.exists_closed_singleton`: Given a closed set `S` in a compact T₀ space, there is some `x ∈ S` such that `{x}` is closed. -* `exists_open_singleton_of_open_finset` Given an open `Finset` `S` in a T₀ space, +* `exists_open_singleton_of_open_finite`: Given an open finite set `S` in a T₀ space, there is some `x ∈ S` such that `{x}` is open. ### T₁ spaces @@ -310,7 +310,7 @@ theorem minimal_nonempty_open_eq_singleton [T0Space X] {s : Set X} (hs : IsOpen #align minimal_nonempty_open_eq_singleton minimal_nonempty_open_eq_singleton /-- Given an open finite set `S` in a T₀ space, there is some `x ∈ S` such that `{x}` is open. -/ -theorem exists_open_singleton_of_open_finite [T0Space X] {s : Set X} (hfin : s.Finite) +theorem exists_isOpen_singleton_of_isOpen_finite [T0Space X] {s : Set X} (hfin : s.Finite) (hne : s.Nonempty) (ho : IsOpen s) : ∃ x ∈ s, IsOpen ({x} : Set X) := by lift s to Finset X using hfin induction' s using Finset.strongInductionOn with s ihs @@ -325,11 +325,12 @@ theorem exists_open_singleton_of_open_finite [T0Space X] {s : Set X} (hfin : s.F refine' fun t hts htne hto => of_not_not fun hts' => ht _ lift t to Finset X using s.finite_toSet.subset hts exact ⟨t, ssubset_iff_subset_ne.2 ⟨hts, mt Finset.coe_inj.2 hts'⟩, htne, hto⟩ -#align exists_open_singleton_of_open_finite exists_open_singleton_of_open_finite +#align exists_open_singleton_of_open_finite exists_isOpen_singleton_of_isOpen_finite theorem exists_open_singleton_of_finite [T0Space X] [Finite X] [Nonempty X] : ∃ x : X, IsOpen ({x} : Set X) := - let ⟨x, _, h⟩ := exists_open_singleton_of_open_finite (Set.toFinite _) univ_nonempty isOpen_univ + let ⟨x, _, h⟩ := exists_isOpen_singleton_of_isOpen_finite (Set.toFinite _) + univ_nonempty isOpen_univ ⟨x, h⟩ #align exists_open_singleton_of_fintype exists_open_singleton_of_finite @@ -727,7 +728,7 @@ theorem nhds_le_nhdsSet_iff [T1Space X] {s : Set X} {x : X} : 𝓝 x ≤ 𝓝ˢ /-- Removing a non-isolated point from a dense set, one still obtains a dense set. -/ theorem Dense.diff_singleton [T1Space X] {s : Set X} (hs : Dense s) (x : X) [NeBot (𝓝[≠] x)] : Dense (s \ {x}) := - hs.inter_of_open_right (dense_compl_singleton x) isOpen_compl_singleton + hs.inter_of_isOpen_right (dense_compl_singleton x) isOpen_compl_singleton #align dense.diff_singleton Dense.diff_singleton /-- Removing a finset from a dense set in a space without isolated points, one still @@ -2157,21 +2158,21 @@ theorem nhds_basis_clopen (x : X) : (𝓝 x).HasBasis (fun s : Set X => x ∈ s exact ⟨V, hUV, V_op, hxV⟩⟩ #align nhds_basis_clopen nhds_basis_clopen -theorem isTopologicalBasis_clopen : IsTopologicalBasis { s : Set X | IsClopen s } := by - apply isTopologicalBasis_of_open_of_nhds fun U (hU : IsClopen U) => hU.1 +theorem isTopologicalBasis_isClopen : IsTopologicalBasis { s : Set X | IsClopen s } := by + apply isTopologicalBasis_of_isOpen_of_nhds fun U (hU : IsClopen U) => hU.1 intro x U hxU U_op have : U ∈ 𝓝 x := IsOpen.mem_nhds U_op hxU rcases (nhds_basis_clopen x).mem_iff.mp this with ⟨V, ⟨hxV, hV⟩, hVU : V ⊆ U⟩ use V tauto -#align is_topological_basis_clopen isTopologicalBasis_clopen +#align is_topological_basis_clopen isTopologicalBasis_isClopen /-- Every member of an open set in a compact Hausdorff totally disconnected space is contained in a clopen set contained in the open set. -/ -theorem compact_exists_clopen_in_open {x : X} {U : Set X} (is_open : IsOpen U) (memU : x ∈ U) : +theorem compact_exists_isClopen_in_isOpen {x : X} {U : Set X} (is_open : IsOpen U) (memU : x ∈ U) : ∃ V : Set X, IsClopen V ∧ x ∈ V ∧ V ⊆ U := - isTopologicalBasis_clopen.mem_nhds_iff.1 (is_open.mem_nhds memU) -#align compact_exists_clopen_in_open compact_exists_clopen_in_open + isTopologicalBasis_isClopen.mem_nhds_iff.1 (is_open.mem_nhds memU) +#align compact_exists_clopen_in_open compact_exists_isClopen_in_isOpen end Profinite @@ -2182,15 +2183,15 @@ variable {H : Type*} [TopologicalSpace H] [LocallyCompactSpace H] [T2Space H] /-- A locally compact Hausdorff totally disconnected space has a basis with clopen elements. -/ theorem loc_compact_Haus_tot_disc_of_zero_dim [TotallyDisconnectedSpace H] : IsTopologicalBasis { s : Set H | IsClopen s } := by - refine isTopologicalBasis_of_open_of_nhds (fun u hu => hu.1) fun x U memU hU => ?_ + refine isTopologicalBasis_of_isOpen_of_nhds (fun u hu => hu.1) fun x U memU hU => ?_ obtain ⟨s, comp, xs, sU⟩ := exists_compact_subset hU memU let u : Set s := ((↑) : s → H) ⁻¹' interior s have u_open_in_s : IsOpen u := isOpen_interior.preimage continuous_subtype_val lift x to s using interior_subset xs haveI : CompactSpace s := isCompact_iff_compactSpace.1 comp - obtain ⟨V : Set s, clopen_in_s, Vx, V_sub⟩ := compact_exists_clopen_in_open u_open_in_s xs - have V_clopen : IsClopen (((↑) : s → H) '' V) := by - refine' ⟨_, comp.isClosed.closedEmbedding_subtype_val.closed_iff_image_closed.1 clopen_in_s.2⟩ + obtain ⟨V : Set s, VisClopen, Vx, V_sub⟩ := compact_exists_isClopen_in_isOpen u_open_in_s xs + have VisClopen' : IsClopen (((↑) : s → H) '' V) := by + refine' ⟨_, comp.isClosed.closedEmbedding_subtype_val.closed_iff_image_closed.1 VisClopen.2⟩ let v : Set u := ((↑) : u → s) ⁻¹' V have : ((↑) : u → H) = ((↑) : s → H) ∘ ((↑) : u → s) := rfl have f0 : Embedding ((↑) : u → H) := embedding_subtype_val.comp embedding_subtype_val @@ -2201,12 +2202,12 @@ theorem loc_compact_Haus_tot_disc_of_zero_dim [TotallyDisconnectedSpace H] : apply Set.inter_eq_self_of_subset_left interior_subset rw [this] apply isOpen_interior - have f2 : IsOpen v := clopen_in_s.1.preimage continuous_subtype_val + have f2 : IsOpen v := VisClopen.1.preimage continuous_subtype_val have f3 : ((↑) : s → H) '' V = ((↑) : u → H) '' v := by rw [this, image_comp, Subtype.image_preimage_coe, inter_eq_self_of_subset_left V_sub] rw [f3] apply f1.isOpenMap v f2 - refine' ⟨(↑) '' V, V_clopen, by simp [Vx], Subset.trans _ sU⟩ + refine' ⟨(↑) '' V, VisClopen', by simp [Vx], Subset.trans _ sU⟩ simp set_option linter.uppercaseLean3 false in #align loc_compact_Haus_tot_disc_of_zero_dim loc_compact_Haus_tot_disc_of_zero_dim diff --git a/Mathlib/Topology/Separation/NotNormal.lean b/Mathlib/Topology/Separation/NotNormal.lean index 0dc73955ff9bc..716dbf2c96da0 100644 --- a/Mathlib/Topology/Separation/NotNormal.lean +++ b/Mathlib/Topology/Separation/NotNormal.lean @@ -26,7 +26,7 @@ https://en.wikipedia.org/wiki/Moore_plane#Proof_that_the_Moore_plane_is_not_norm theorem IsClosed.mk_lt_continuum [NormalSpace X] {s : Set X} (hs : IsClosed s) [DiscreteTopology s] : #s < 𝔠 := by -- Proof by contradiction: assume `𝔠 ≤ #s` - by_contra' h + by_contra! h -- Choose a countable dense set `t : Set X` rcases exists_countable_dense X with ⟨t, htc, htd⟩ haveI := htc.to_subtype diff --git a/Mathlib/Topology/Sets/Opens.lean b/Mathlib/Topology/Sets/Opens.lean index 6a85c457357fe..50568cf62e67b 100644 --- a/Mathlib/Topology/Sets/Opens.lean +++ b/Mathlib/Topology/Sets/Opens.lean @@ -302,7 +302,7 @@ theorem isBasis_iff_nbhd {B : Set (Opens α)} : dsimp at H₂ subst H₂ exact hsV - · refine' isTopologicalBasis_of_open_of_nhds _ _ + · refine' isTopologicalBasis_of_isOpen_of_nhds _ _ · rintro sU ⟨U, -, rfl⟩ exact U.2 · intro x sU hx hsU diff --git a/Mathlib/Topology/Sheaves/Operations.lean b/Mathlib/Topology/Sheaves/Operations.lean index acd5c539ba110..1c04fe1fd7316 100644 --- a/Mathlib/Topology/Sheaves/Operations.lean +++ b/Mathlib/Topology/Sheaves/Operations.lean @@ -134,9 +134,9 @@ instance (F : X.Sheaf CommRingCat.{w}) : Mono F.presheaf.toTotalQuotientPresheaf intro x rw [map_zero] apply Submonoid.mem_iInf.mp hs x - -- Porting note : added `dsimp` to make `rw ←map_mul` work + -- Porting note : added `dsimp` to make `rw [← map_mul]` work dsimp - rw [←map_mul, e, map_zero] + rw [← map_mul, e, map_zero] end Presheaf diff --git a/Mathlib/Topology/Sheaves/Presheaf.lean b/Mathlib/Topology/Sheaves/Presheaf.lean index 9fc3a4158b33e..a4831770ba571 100644 --- a/Mathlib/Topology/Sheaves/Presheaf.lean +++ b/Mathlib/Topology/Sheaves/Presheaf.lean @@ -454,7 +454,7 @@ theorem toPushforwardOfIso_app {X Y : TopCat} (H₁ : X ≅ Y) {ℱ : X.Presheaf Functor.id_obj, Functor.comp_obj, Iso.symm_hom, NatIso.op_inv, Iso.symm_inv, NatTrans.op_app, NatIso.ofComponents_hom_app, eqToIso.hom, eqToHom_op, Equivalence.Equivalence_mk'_unitInv, Equivalence.Equivalence_mk'_counitInv, NatIso.op_hom, unop_op, op_unop, eqToIso.inv, - NatIso.ofComponents_inv_app, eqToHom_unop, ←ℱ.map_comp, eqToHom_trans, eqToHom_map, + NatIso.ofComponents_inv_app, eqToHom_unop, ← ℱ.map_comp, eqToHom_trans, eqToHom_map, presheafEquivOfIso_unitIso_hom_app_app] set_option linter.uppercaseLean3 false in #align Top.presheaf.to_pushforward_of_iso_app TopCat.Presheaf.toPushforwardOfIso_app diff --git a/Mathlib/Topology/Sheaves/SheafCondition/PairwiseIntersections.lean b/Mathlib/Topology/Sheaves/SheafCondition/PairwiseIntersections.lean index bf532db9d0719..d60f523ea79b1 100644 --- a/Mathlib/Topology/Sheaves/SheafCondition/PairwiseIntersections.lean +++ b/Mathlib/Topology/Sheaves/SheafCondition/PairwiseIntersections.lean @@ -375,7 +375,7 @@ def interUnionPullbackConeLift : s.pt ⟶ F.1.obj (op (U ⊔ V)) := by dsimp [Pairwise.diagram] <;> simp only [Category.id_comp, s.condition, CategoryTheory.Functor.map_id, Category.comp_id] · rw [← cancel_mono (F.1.map (eqToHom <| inf_comm : U ⊓ V ⟶ _).op), Category.assoc, - Category.assoc, ←F.1.map_comp, ←F.1.map_comp] + Category.assoc, ← F.1.map_comp, ← F.1.map_comp] exact s.condition.symm set_option linter.uppercaseLean3 false in #align Top.sheaf.inter_union_pullback_cone_lift TopCat.Sheaf.interUnionPullbackConeLift diff --git a/Mathlib/Topology/Sheaves/Stalks.lean b/Mathlib/Topology/Sheaves/Stalks.lean index 0eb22f953cd99..681d9633195c9 100644 --- a/Mathlib/Topology/Sheaves/Stalks.lean +++ b/Mathlib/Topology/Sheaves/Stalks.lean @@ -112,7 +112,7 @@ set_option linter.uppercaseLean3 false in -- Porting note : `@[elementwise]` did not generate the best lemma when applied to `germ_res` theorem germ_res_apply (F : X.Presheaf C) {U V : Opens X} (i : U ⟶ V) (x : U) [ConcreteCategory C] - (s) : germ F x (F.map i.op s) = germ F (i x) s := by rw [←comp_apply, germ_res] + (s) : germ F x (F.map i.op s) = germ F (i x) s := by rw [← comp_apply, germ_res] set_option linter.uppercaseLean3 false in #align Top.presheaf.germ_res_apply TopCat.Presheaf.germ_res_apply diff --git a/Mathlib/Topology/UniformSpace/AbstractCompletion.lean b/Mathlib/Topology/UniformSpace/AbstractCompletion.lean index afc421f38e32e..425558ebac6c4 100644 --- a/Mathlib/Topology/UniformSpace/AbstractCompletion.lean +++ b/Mathlib/Topology/UniformSpace/AbstractCompletion.lean @@ -208,7 +208,7 @@ theorem map_unique {f : α → β} {g : hatα → hatβ} (hg : UniformContinuous pkg.funext (pkg.continuous_map _ _) hg.continuous <| by intro a change pkg.extend (ι' ∘ f) _ = _ - simp_rw [(· ∘ ·), h, ←comp_apply (f := g)] + simp_rw [(· ∘ ·), h, ← comp_apply (f := g)] rw [pkg.extend_coe (hg.comp pkg.uniformContinuous_coe)] #align abstract_completion.map_unique AbstractCompletion.map_unique diff --git a/Mathlib/Topology/UniformSpace/Basic.lean b/Mathlib/Topology/UniformSpace/Basic.lean index 434fe786b76b8..c253e00725bea 100644 --- a/Mathlib/Topology/UniformSpace/Basic.lean +++ b/Mathlib/Topology/UniformSpace/Basic.lean @@ -1778,8 +1778,8 @@ theorem union_mem_uniformity_sum {a : Set (α × α)} (ha : a ∈ 𝓤 α) {b : /- To prove that the topology defined by the uniform structure on the disjoint union coincides with the disjoint union topology, we need two lemmas saying that open sets can be characterized by the uniform structure -/ -theorem uniformity_sum_of_open_aux {s : Set (Sum α β)} (hs : IsOpen s) {x : Sum α β} (xs : x ∈ s) : - { p : (α ⊕ β) × (α ⊕ β) | p.1 = x → p.2 ∈ s } ∈ +theorem uniformity_sum_of_isOpen_aux {s : Set (Sum α β)} (hs : IsOpen s) {x : Sum α β} + (xs : x ∈ s) : { p : (α ⊕ β) × (α ⊕ β) | p.1 = x → p.2 ∈ s } ∈ (@UniformSpace.Core.sum α β _ _).uniformity := by cases x · refine' mem_of_superset @@ -1790,9 +1790,9 @@ theorem uniformity_sum_of_open_aux {s : Set (Sum α β)} (hs : IsOpen s) {x : Su (union_mem_uniformity_sum univ_mem (mem_nhds_uniformity_iff_right.1 (hs.2.mem_nhds xs))) (union_subset _ _) <;> rintro _ ⟨⟨a, _⟩, h, ⟨⟩⟩ ⟨⟩ exact h rfl -#align uniformity_sum_of_open_aux uniformity_sum_of_open_aux +#align uniformity_sum_of_open_aux uniformity_sum_of_isOpen_aux -theorem open_of_uniformity_sum_aux {s : Set (Sum α β)} +theorem isOpen_of_uniformity_sum_aux {s : Set (Sum α β)} (hs : ∀ x ∈ s, { p : (α ⊕ β) × (α ⊕ β) | p.1 = x → p.2 ∈ s } ∈ (@UniformSpace.Core.sum α β _ _).uniformity) : IsOpen s := by @@ -1807,12 +1807,12 @@ theorem open_of_uniformity_sum_aux {s : Set (Sum α β)} refine' mem_of_superset ht _ rintro p pt rfl exact st ⟨_, pt, rfl⟩ rfl -#align open_of_uniformity_sum_aux open_of_uniformity_sum_aux +#align open_of_uniformity_sum_aux isOpen_of_uniformity_sum_aux -- We can now define the uniform structure on the disjoint union instance Sum.uniformSpace : UniformSpace (Sum α β) where toCore := UniformSpace.Core.sum - isOpen_uniformity _ := ⟨uniformity_sum_of_open_aux, open_of_uniformity_sum_aux⟩ + isOpen_uniformity _ := ⟨uniformity_sum_of_isOpen_aux, isOpen_of_uniformity_sum_aux⟩ #align sum.uniform_space Sum.uniformSpace theorem Sum.uniformity : diff --git a/Mathlib/Topology/UniformSpace/Cauchy.lean b/Mathlib/Topology/UniformSpace/Cauchy.lean index 08d753da49dd0..322c3a0b9df97 100644 --- a/Mathlib/Topology/UniformSpace/Cauchy.lean +++ b/Mathlib/Topology/UniformSpace/Cauchy.lean @@ -799,7 +799,7 @@ theorem secondCountable_of_separable [SeparableSpace α] : SecondCountableTopolo (@uniformity_hasBasis_open_symmetric α _).exists_antitone_subbasis choose ht_mem hto hts using hto refine' ⟨⟨⋃ x ∈ s, range fun k => ball x (t k), hsc.biUnion fun x _ => countable_range _, _⟩⟩ - refine' (isTopologicalBasis_of_open_of_nhds _ _).eq_generateFrom + refine' (isTopologicalBasis_of_isOpen_of_nhds _ _).eq_generateFrom · simp only [mem_iUnion₂, mem_range] rintro _ ⟨x, _, k, rfl⟩ exact isOpen_ball x (hto k) diff --git a/Mathlib/Topology/UnitInterval.lean b/Mathlib/Topology/UnitInterval.lean index 4ab4fa3ceaac0..62fcbbe9a50d0 100644 --- a/Mathlib/Topology/UnitInterval.lean +++ b/Mathlib/Topology/UnitInterval.lean @@ -123,6 +123,9 @@ theorem symm_symm (x : I) : σ (σ x) = x := Subtype.ext <| by simp [symm] #align unit_interval.symm_symm unitInterval.symm_symm +theorem symm_bijective : Function.Bijective (symm : I → I) := + Function.bijective_iff_has_inverse.mpr ⟨_, symm_symm, symm_symm⟩ + @[simp] theorem coe_symm_eq (x : I) : (σ x : ℝ) = 1 - x := rfl @@ -141,7 +144,7 @@ theorem involutive_symm : Function.Involutive σ := symm_symm theorem bijective_symm : Function.Bijective σ := involutive_symm.bijective theorem half_le_symm_iff (t : I) : 1 / 2 ≤ (σ t : ℝ) ↔ (t : ℝ) ≤ 1 / 2 := by - rw [coe_symm_eq, le_sub_iff_add_le, add_comm, ←le_sub_iff_add_le, sub_half] + rw [coe_symm_eq, le_sub_iff_add_le, add_comm, ← le_sub_iff_add_le, sub_half] instance : ConnectedSpace I := Subtype.connectedSpace ⟨nonempty_Icc.mpr zero_le_one, isPreconnected_Icc⟩ @@ -177,7 +180,7 @@ theorem le_one' {t : I} : t ≤ 1 := t.2.2 #align unit_interval.le_one' unitInterval.le_one' -instance : NeZero (1 : I) := ⟨fun h ↦ one_ne_zero <| congrArg Subtype.val h⟩ +instance : Nontrivial I := ⟨⟨1, 0, (one_ne_zero <| congrArg Subtype.val ·)⟩⟩ theorem mul_pos_mem_iff {a t : ℝ} (ha : 0 < a) : a * t ∈ I ↔ t ∈ Set.Icc (0 : ℝ) (1 / a) := by constructor <;> rintro ⟨h₁, h₂⟩ <;> constructor diff --git a/docs/overview.yaml b/docs/overview.yaml index b87559d193664..d262da1176d36 100644 --- a/docs/overview.yaml +++ b/docs/overview.yaml @@ -249,7 +249,7 @@ Topology: Lipschitz continuity: 'LipschitzWith' Hölder continuity: 'HolderWith' contraction mapping theorem: 'ContractingWith.exists_fixedPoint' - Baire theorem: 'dense_iInter_of_open' + Baire theorem: 'dense_iInter_of_isOpen' Arzela-Ascoli theorem: 'BoundedContinuousFunction.arzela_ascoli' Hausdorff distance: 'EMetric.hausdorffEdist' Gromov-Hausdorff space: 'GromovHausdorff.GHSpace' diff --git a/lake-manifest.json b/lake-manifest.json index fad26333341b3..b75620c95831e 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -4,7 +4,7 @@ [{"url": "https://github.com/leanprover/std4", "type": "git", "subDir": null, - "rev": "8f543b17ad95f3741d5e449723466f1f3d02bc8f", + "rev": "bc0072ecff2275e623f938d648fadabfc7980147", "name": "std", "manifestFile": "lake-manifest.json", "inputRev": "nightly-testing", @@ -13,7 +13,7 @@ {"url": "https://github.com/leanprover-community/quote4", "type": "git", "subDir": null, - "rev": "d3a1d25f3eba0d93a58d5d3d027ffa78ece07755", + "rev": "ccba5d35d07a448fab14c0e391c8105df6e2564c", "name": "Qq", "manifestFile": "lake-manifest.json", "inputRev": "master", @@ -22,7 +22,7 @@ {"url": "https://github.com/leanprover-community/aesop", "type": "git", "subDir": null, - "rev": "c7cff4551258d31c0d2d453b3f9cbca757d445f1", + "rev": "3141402ba5a5f0372d2378fd75a481bc79a74ecf", "name": "aesop", "manifestFile": "lake-manifest.json", "inputRev": "master", diff --git a/lakefile.lean b/lakefile.lean index 9c83f7364ac35..735fa58298481 100644 --- a/lakefile.lean +++ b/lakefile.lean @@ -2,30 +2,22 @@ import Lake open Lake DSL -def moreServerArgs := #[ - "-Dpp.unicode.fun=true", -- pretty-prints `fun a ↦ b` - "-Dpp.proofs.withType=false", - "-DautoImplicit=false", - "-DrelaxedAutoImplicit=false" -] - --- These settings only apply during `lake build`, but not in VSCode editor. -def moreLeanArgs := moreServerArgs - --- These are additional settings which do not affect the lake hash, --- so they can be enabled in CI and disabled locally or vice versa. --- Warning: Do not put any options here that actually change the olean files, --- or inconsistent behavior may result -def weakLeanArgs : Array String := - if get_config? CI |>.isSome then - #["-DwarningAsError=true"] - else - #[] - package mathlib where - moreServerArgs := moreServerArgs - moreLeanArgs := moreLeanArgs - weakLeanArgs := weakLeanArgs + leanOptions := #[ + ⟨`pp.unicode.fun, true⟩, -- pretty-prints `fun a ↦ b` + ⟨`pp.proofs.withType, false⟩, + ⟨`autoImplicit, false⟩, + ⟨`relaxedAutoImplicit, false⟩ + ] + -- These are additional settings which do not affect the lake hash, + -- so they can be enabled in CI and disabled locally or vice versa. + -- Warning: Do not put any options here that actually change the olean files, + -- or inconsistent behavior may result + weakLeanArgs := + if get_config? CI |>.isSome then + #["-DwarningAsError=true"] + else + #[] /-! ## Mathlib dependencies on upstream projects. diff --git a/scripts/bench/temci-config.run.yml b/scripts/bench/temci-config.run.yml index d73e75ff5569c..b737b9776c1c4 100644 --- a/scripts/bench/temci-config.run.yml +++ b/scripts/bench/temci-config.run.yml @@ -26,7 +26,8 @@ properties: ['wall-clock', 'task-clock', 'instructions'] rusage_properties: ['maxrss'] cmd: | - LEAN_PATH=$(lake print-paths --no-build Mathlib | jq -r '.oleanPath | join(":")') lean Mathlib.lean + # run lake+lean like the file worker would + LEAN_PATH=$(lake setup-file --no-build Mathlib Mathlib | jq -r '.paths.oleanPath | join(":")') lean Mathlib.lean discarded_runs: 1 min_runs: 5 - attributes: diff --git a/scripts/lean-pr-testing-comments.sh b/scripts/lean-pr-testing-comments.sh index 9d19e59ce7469..854c7d6f25a72 100755 --- a/scripts/lean-pr-testing-comments.sh +++ b/scripts/lean-pr-testing-comments.sh @@ -67,7 +67,7 @@ if [[ "$branch_name" =~ ^lean-pr-testing-([0-9]+)$ ]]; then existing_comment=$(curl -L -s -H "Authorization: token $TOKEN" \ -H "Accept: application/vnd.github.v3+json" \ "https://api.github.com/repos/leanprover/lean4/issues/$pr_number/comments" \ - | jq '.[] | select(.body | startswith("- ❗ Mathlib") or startswith("- ✅ Mathlib") or startswith("- ❌ Mathlib") or startswith("- 💥 Mathlib") or startswith("- 🟡 Mathlib"))') + | jq '.[] | select(.body | test("^- . Mathlib")) | select(.user.login == "leanprover-community-mathlib4-bot")') existing_comment_id=$(echo "$existing_comment" | jq -r .id) existing_comment_body=$(echo "$existing_comment" | jq -r .body) diff --git a/test/LibrarySearch/basic.lean b/test/LibrarySearch/basic.lean index 08fa256d5850f..5a101cdd77164 100644 --- a/test/LibrarySearch/basic.lean +++ b/test/LibrarySearch/basic.lean @@ -223,10 +223,16 @@ lemma prime_of_prime (n : ℕ) : Prime n ↔ Nat.Prime n := by lemma ex' (x : ℕ) (_h₁ : x = 0) (h : 2 * 2 ∣ x) : 2 ∣ x := by exact? says exact dvd_of_mul_left_dvd h --- Example from https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/Exact.3F.20fails.20on.20le_antisymm/near/388993167 -example {x y : ℝ} (hxy : x ≤ y) (hyx : y ≤ x) : x = y := by - exact? says exact le_antisymm hxy hyx - -- https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/apply.3F.20failure/near/402534407 example (P Q : Prop) (h : P → Q) (h' : ¬Q) : ¬P := by exact? says exact mt h h' + +-- Removed until we come up with a way of handling nonspecific lemmas +-- that does not pollute the output or cause too much slow-down. +-- -- Example from https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/Exact.3F.20fails.20on.20le_antisymm/near/388993167 +-- set_option linter.unreachableTactic false in +-- example {x y : ℝ} (hxy : x ≤ y) (hyx : y ≤ x) : x = y := by +-- -- This example non-deterministically picks between `le_antisymm hxy hyx` and `ge_antisymm hyx hxy`. +-- first +-- | exact? says exact le_antisymm hxy hyx +-- | exact? says exact ge_antisymm hyx hxy diff --git a/test/NthRewrite.lean b/test/NthRewrite.lean index 3838b74c45ded..09f122a455847 100644 --- a/test/NthRewrite.lean +++ b/test/NthRewrite.lean @@ -6,11 +6,11 @@ import Mathlib.Data.Nat.Basic set_option autoImplicit true example [AddZeroClass G] {a : G} (h : a = a): a = (a + 0) := by - nth_rewrite 2 [←add_zero a] at h + nth_rewrite 2 [← add_zero a] at h exact h example [AddZeroClass G] {a : G} : a + a = a + (a + 0) := by - nth_rw 2 [←add_zero a] + nth_rw 2 [← add_zero a] structure F := (a : ℕ) @@ -57,8 +57,8 @@ axiom bar' : [[5],[5]] = [[6],[6]] example : [[7],[6]] = [[5],[5]] := by nth_rewrite 1 [foo'] nth_rewrite 1 [bar'] - nth_rewrite 1 [←foo'] - nth_rewrite 1 [←foo'] + nth_rewrite 1 [← foo'] + nth_rewrite 1 [← foo'] rfl -- Porting note: @@ -80,7 +80,7 @@ example : [[7],[6]] = [[5],[5]] := by -- example : [(3, 3), (5, 9), (5, 9)] = [(4, 5), (3, 6), (1, 1)] := by -- nth_rewrite 1 [wowzer] --- nth_rewrite 3 [←pchew] +-- nth_rewrite 3 [← pchew] -- nth_rewrite 1 [pchew] -- nth_rewrite 1 [smash] diff --git a/test/byContra.lean b/test/byContra.lean index 04e92a31db023..6cef8ec43c7b5 100644 --- a/test/byContra.lean +++ b/test/byContra.lean @@ -1,4 +1,4 @@ --- tests for byContra' tactic +-- tests for by_contra! tactic import Mathlib.Tactic.ByContra import Mathlib.Tactic.Rename import Mathlib.Tactic.Set @@ -6,32 +6,32 @@ import Mathlib.Data.Nat.Basic set_option autoImplicit true example (a b : ℕ) (foo : False) : a < b := by - by_contra' + by_contra! guard_hyp this : b ≤ a exact foo example (a b : ℕ) (h : False) : a < b := by - by_contra' foo + by_contra! foo revert foo; change b ≤ a → False; intro; exact h example (a b : ℕ) (h : False) : a < b := by - by_contra' foo : ¬ a < b -- can avoid push_neg + by_contra! foo : ¬ a < b -- can avoid push_neg guard_hyp foo : ¬ a < b exact h example : 1 < 2 := by - by_contra' + by_contra! guard_hyp this : 2 ≤ 1 contradiction example (p : Prop) (bar : False) : ¬ ¬ ¬ ¬ ¬ ¬ P := by - by_contra' foo : ¬ ¬ ¬ P -- normalises to ¬ P, as does ¬ (goal). + by_contra! foo : ¬ ¬ ¬ P -- normalises to ¬ P, as does ¬ (goal). guard_hyp foo : ¬ ¬ ¬ P exact bar example (p : Prop) (bar : False) : ¬ ¬ ¬ ¬ ¬ ¬ P := by - by_contra' : ¬ ¬ ¬ P + by_contra! : ¬ ¬ ¬ P guard_hyp this : ¬ ¬ ¬ P exact bar @@ -39,19 +39,19 @@ variable [LinearOrder α] [One α] [Mul α] example (x : α) (f : False) : x ≤ 1 := by set a := x * x - by_contra' h + by_contra! h guard_hyp h : 1 < x assumption example (x : α) (f : False) : x ≤ 1 := by let _a := x * x - by_contra' h + by_contra! h guard_hyp h : 1 < x assumption example (x : α) (f : False) : x ≤ 1 := by set a := x * x have : a ≤ a := le_rfl - by_contra' h + by_contra! h guard_hyp h : 1 < x assumption diff --git a/test/choose_reduction.lean b/test/choose_reduction.lean new file mode 100644 index 0000000000000..ac200c8b78703 --- /dev/null +++ b/test/choose_reduction.lean @@ -0,0 +1,49 @@ +import Mathlib.Data.Finset.Basic +/-! +Tests that `List.choose` and friends all reduce appropriately. + +Previously this was blocked by reducibility issues with `And.rec` +(https://leanprover.zulipchat.com/#narrow/stream/236446-Type-theory/topic/And.2Erec/near/398483665). +-/ + +namespace List + +lemma foo : ∃ x, x ∈ ([0,3] : List (Fin 5)) ∧ x = 3 := by + use 3 + simp + +example : choose _ _ foo = 3 := by rfl +example : choose _ _ foo = 3 := rfl +example : choose _ _ foo = 3 := by eq_refl +example : choose _ _ foo = 3 := by exact rfl +example : choose _ _ foo = 3 := by decide + +end List + +namespace Multiset + +lemma foo : ∃! x, x ∈ ({0,3} : Multiset (Fin 5)) ∧ x = 3 := by + use 3 + simp + +example : choose _ _ foo = 3 := by rfl +example : choose _ _ foo = 3 := rfl +example : choose _ _ foo = 3 := by eq_refl +example : choose _ _ foo = 3 := by exact rfl +example : choose _ _ foo = 3 := by decide + +end Multiset + +namespace Finset + +lemma foo : ∃! x, x ∈ ({0,3} : Finset (Fin 5)) ∧ x = 3 := by + use 3 + simp + +example : choose _ _ foo = 3 := by rfl +example : choose _ _ foo = 3 := rfl +example : choose _ _ foo = 3 := by eq_refl +example : choose _ _ foo = 3 := by exact rfl +example : choose _ _ foo = 3 := by decide + +end Finset diff --git a/test/convert.lean b/test/convert.lean index 0e4db7387a504..2394b179c4f40 100644 --- a/test/convert.lean +++ b/test/convert.lean @@ -38,7 +38,7 @@ example (α β : Type) (h : α = β) (b : β) : Nat × Nat × Nat × α := by example {f : β → α} {x y : α} (h : x ≠ y) : f ⁻¹' {x} ∩ f ⁻¹' {y} = ∅ := by have : {x} ∩ {y} = (∅ : Set α) := by simpa [ne_comm] using h convert Set.preimage_empty - rw [←Set.preimage_inter, this] + rw [← Set.preimage_inter, this] section convert_to diff --git a/test/linarith.lean b/test/linarith.lean index 96d994e6c21f2..cb89b23fcd5f0 100644 --- a/test/linarith.lean +++ b/test/linarith.lean @@ -490,8 +490,7 @@ end -- Checks that splitNe handles metavariables and also that conjunction splitting occurs -- before splitNe splitting example (r : ℚ) (h' : 1 = r * 2) : 1 = 0 ∨ r = 1 / 2 := by - by_contra h'' - push_neg at h'' + by_contra! h'' linarith (config := {splitNe := true}) example (x y : ℚ) (h₁ : 0 ≤ y) (h₂ : y ≤ x) : y * x ≤ x * x := by nlinarith diff --git a/test/observe.lean b/test/observe.lean index 0cef4794d77c8..24d4b0d6252b9 100644 --- a/test/observe.lean +++ b/test/observe.lean @@ -13,7 +13,7 @@ theorem euclid (n : ℕ) : ∃ N, n < N ∧ N.Prime := by observe : n.factorial > 0 linarith constructor - · by_contra' + · by_contra! observe : p ∣ n.factorial observe : p ∣ N observe : p ∣ 1