diff --git a/.github/build.in.yml b/.github/build.in.yml index b7acc84572050..69c40f7765dae 100644 --- a/.github/build.in.yml +++ b/.github/build.in.yml @@ -15,7 +15,8 @@ env: concurrency: # label each workflow run; only the latest with each label will run # workflows on master get more expressive labels - group: ${{ github.workflow }}-${{ github.ref }}.${{(github.ref == 'refs/heads/master' && github.run_id) || ''}} + group: ${{ github.workflow }}-${{ github.ref }}. + ${{ ( contains(fromJSON( '["refs/heads/master", "refs/heads/staging"]'), github.ref ) && github.run_id) || ''}} # cancel any running workflow with the same label cancel-in-progress: true diff --git a/.github/workflows/bors.yml b/.github/workflows/bors.yml index 37c0ac50d19a0..3149eccc52722 100644 --- a/.github/workflows/bors.yml +++ b/.github/workflows/bors.yml @@ -25,7 +25,8 @@ env: concurrency: # label each workflow run; only the latest with each label will run # workflows on master get more expressive labels - group: ${{ github.workflow }}-${{ github.ref }}.${{(github.ref == 'refs/heads/master' && github.run_id) || ''}} + group: ${{ github.workflow }}-${{ github.ref }}. + ${{ ( contains(fromJSON( '["refs/heads/master", "refs/heads/staging"]'), github.ref ) && github.run_id) || ''}} # cancel any running workflow with the same label cancel-in-progress: true diff --git a/.github/workflows/build.yml b/.github/workflows/build.yml index 62525af3c3089..5a3bc5d84c1b0 100644 --- a/.github/workflows/build.yml +++ b/.github/workflows/build.yml @@ -32,7 +32,8 @@ env: concurrency: # label each workflow run; only the latest with each label will run # workflows on master get more expressive labels - group: ${{ github.workflow }}-${{ github.ref }}.${{(github.ref == 'refs/heads/master' && github.run_id) || ''}} + group: ${{ github.workflow }}-${{ github.ref }}. + ${{ ( contains(fromJSON( '["refs/heads/master", "refs/heads/staging"]'), github.ref ) && github.run_id) || ''}} # cancel any running workflow with the same label cancel-in-progress: true diff --git a/.github/workflows/build_fork.yml b/.github/workflows/build_fork.yml index 14ae2662c9d09..d7ceebe6b73e3 100644 --- a/.github/workflows/build_fork.yml +++ b/.github/workflows/build_fork.yml @@ -29,7 +29,8 @@ env: concurrency: # label each workflow run; only the latest with each label will run # workflows on master get more expressive labels - group: ${{ github.workflow }}-${{ github.ref }}.${{(github.ref == 'refs/heads/master' && github.run_id) || ''}} + group: ${{ github.workflow }}-${{ github.ref }}. + ${{ ( contains(fromJSON( '["refs/heads/master", "refs/heads/staging"]'), github.ref ) && github.run_id) || ''}} # cancel any running workflow with the same label cancel-in-progress: true diff --git a/Archive/Sensitivity.lean b/Archive/Sensitivity.lean index ab52f935c6633..7f94a65f75cbd 100644 --- a/Archive/Sensitivity.lean +++ b/Archive/Sensitivity.lean @@ -368,7 +368,7 @@ open Classical in subspace of `V (m+1)` spanned by the corresponding basis vectors non-trivially intersects the range of `g m`. -/ theorem exists_eigenvalue (H : Set (Q m.succ)) (hH : Card H ≥ 2 ^ m + 1) : - ∃ y ∈ Span (e '' H) ⊓ range (g m), y ≠ (0 : _) := by + ∃ y ∈ Span (e '' H) ⊓ range (g m), y ≠ 0 := by let W := Span (e '' H) let img := range (g m) suffices 0 < dim (W ⊓ img) by @@ -408,7 +408,7 @@ theorem huang_degree_theorem (H : Set (Q m.succ)) (hH : Card H ≥ 2 ^ m + 1) : rw [Finsupp.mem_support_iff] at p_in rw [Set.mem_toFinset] exact (dualBases_e_ε _).mem_of_mem_span y_mem_H p p_in - obtain ⟨q, H_max⟩ : ∃ q : Q m.succ, ∀ q' : Q m.succ, |(ε q' : _) y| ≤ |ε q y| := + obtain ⟨q, H_max⟩ : ∃ q : Q m.succ, ∀ q' : Q m.succ, |(ε q' :) y| ≤ |ε q y| := Finite.exists_max _ have H_q_pos : 0 < |ε q y| := by contrapose! y_ne diff --git a/Counterexamples/Girard.lean b/Counterexamples/Girard.lean index 2045588c13e53..16ea6890ecf9a 100644 --- a/Counterexamples/Girard.lean +++ b/Counterexamples/Girard.lean @@ -41,7 +41,7 @@ theorem girard.{u} (pi : (Type u → Type u) → Type u) let τ (T : Set (Set U)) : U := lam (G T) let σ (S : U) : Set (Set U) := app S U τ have στ : ∀ {s S}, s ∈ σ (τ S) ↔ {x | τ (σ x) ∈ s} ∈ S := fun {s S} => - iff_of_eq (congr_arg (fun f : F U => s ∈ f τ) (beta (G S) U) : _) + iff_of_eq (congr_arg (fun f : F U => s ∈ f τ) (beta (G S) U) :) let ω : Set (Set U) := {p | ∀ x, p ∈ σ x → x ∈ p} let δ (S : Set (Set U)) := ∀ p, p ∈ S → τ S ∈ p have : δ ω := fun _p d => d (τ ω) <| στ.2 fun x h => d (τ (σ x)) (στ.2 h) diff --git a/Mathlib.lean b/Mathlib.lean index f692eadfdec90..439e51e558829 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -689,6 +689,7 @@ import Mathlib.Algebra.Order.Group.Instances import Mathlib.Algebra.Order.Group.Int import Mathlib.Algebra.Order.Group.Lattice import Mathlib.Algebra.Order.Group.MinMax +import Mathlib.Algebra.Order.Group.Multiset import Mathlib.Algebra.Order.Group.Nat import Mathlib.Algebra.Order.Group.Opposite import Mathlib.Algebra.Order.Group.OrderIso @@ -1081,6 +1082,7 @@ import Mathlib.AlgebraicTopology.SimplicialCategory.SimplicialObject import Mathlib.AlgebraicTopology.SimplicialNerve import Mathlib.AlgebraicTopology.SimplicialObject.Basic import Mathlib.AlgebraicTopology.SimplicialObject.Coskeletal +import Mathlib.AlgebraicTopology.SimplicialObject.Split import Mathlib.AlgebraicTopology.SimplicialSet.Basic import Mathlib.AlgebraicTopology.SimplicialSet.Coskeletal import Mathlib.AlgebraicTopology.SimplicialSet.HomotopyCat @@ -1090,7 +1092,6 @@ import Mathlib.AlgebraicTopology.SimplicialSet.Nerve import Mathlib.AlgebraicTopology.SimplicialSet.Path import Mathlib.AlgebraicTopology.SimplicialSet.StrictSegal import Mathlib.AlgebraicTopology.SingularSet -import Mathlib.AlgebraicTopology.SplitSimplicialObject import Mathlib.AlgebraicTopology.TopologicalSimplex import Mathlib.Analysis.Analytic.Basic import Mathlib.Analysis.Analytic.CPolynomial @@ -2212,6 +2213,7 @@ import Mathlib.CategoryTheory.Subterminal import Mathlib.CategoryTheory.Sums.Associator import Mathlib.CategoryTheory.Sums.Basic import Mathlib.CategoryTheory.Thin +import Mathlib.CategoryTheory.Triangulated.Adjunction import Mathlib.CategoryTheory.Triangulated.Basic import Mathlib.CategoryTheory.Triangulated.Functor import Mathlib.CategoryTheory.Triangulated.HomologicalFunctor @@ -2456,6 +2458,7 @@ import Mathlib.Data.DFinsupp.NeLocus import Mathlib.Data.DFinsupp.Notation import Mathlib.Data.DFinsupp.Order import Mathlib.Data.DFinsupp.Sigma +import Mathlib.Data.DFinsupp.Small import Mathlib.Data.DFinsupp.Submonoid import Mathlib.Data.DFinsupp.WellFounded import Mathlib.Data.DList.Instances @@ -3761,10 +3764,12 @@ import Mathlib.MeasureTheory.MeasurableSpace.EventuallyMeasurable import Mathlib.MeasureTheory.MeasurableSpace.Instances import Mathlib.MeasureTheory.MeasurableSpace.Invariants import Mathlib.MeasureTheory.MeasurableSpace.NCard +import Mathlib.MeasureTheory.MeasurableSpace.Pi import Mathlib.MeasureTheory.MeasurableSpace.PreorderRestrict import Mathlib.MeasureTheory.MeasurableSpace.Prod import Mathlib.MeasureTheory.Measure.AEDisjoint import Mathlib.MeasureTheory.Measure.AEMeasurable +import Mathlib.MeasureTheory.Measure.AbsolutelyContinuous import Mathlib.MeasureTheory.Measure.AddContent import Mathlib.MeasureTheory.Measure.Comap import Mathlib.MeasureTheory.Measure.Complex @@ -3794,6 +3799,7 @@ import Mathlib.MeasureTheory.Measure.Lebesgue.Integral import Mathlib.MeasureTheory.Measure.Lebesgue.VolumeOfBalls import Mathlib.MeasureTheory.Measure.LevyProkhorovMetric import Mathlib.MeasureTheory.Measure.LogLikelihoodRatio +import Mathlib.MeasureTheory.Measure.Map import Mathlib.MeasureTheory.Measure.MeasureSpace import Mathlib.MeasureTheory.Measure.MeasureSpaceDef import Mathlib.MeasureTheory.Measure.MutuallySingular @@ -3802,7 +3808,9 @@ import Mathlib.MeasureTheory.Measure.OpenPos import Mathlib.MeasureTheory.Measure.Portmanteau import Mathlib.MeasureTheory.Measure.ProbabilityMeasure import Mathlib.MeasureTheory.Measure.Prod +import Mathlib.MeasureTheory.Measure.QuasiMeasurePreserving import Mathlib.MeasureTheory.Measure.Regular +import Mathlib.MeasureTheory.Measure.RegularityCompacts import Mathlib.MeasureTheory.Measure.Restrict import Mathlib.MeasureTheory.Measure.SeparableMeasure import Mathlib.MeasureTheory.Measure.Stieltjes @@ -3893,6 +3901,7 @@ import Mathlib.NumberTheory.EulerProduct.ExpLog import Mathlib.NumberTheory.FLT.Basic import Mathlib.NumberTheory.FLT.Four import Mathlib.NumberTheory.FLT.MasonStothers +import Mathlib.NumberTheory.FLT.Polynomial import Mathlib.NumberTheory.FLT.Three import Mathlib.NumberTheory.FactorisationProperties import Mathlib.NumberTheory.Fermat @@ -4387,6 +4396,8 @@ import Mathlib.RingTheory.EisensteinCriterion import Mathlib.RingTheory.EssentialFiniteness import Mathlib.RingTheory.Etale.Basic import Mathlib.RingTheory.Etale.Field +import Mathlib.RingTheory.Etale.Kaehler +import Mathlib.RingTheory.Etale.Pi import Mathlib.RingTheory.EuclideanDomain import Mathlib.RingTheory.Extension import Mathlib.RingTheory.Filtration @@ -4641,6 +4652,7 @@ import Mathlib.RingTheory.RingHom.Integral import Mathlib.RingTheory.RingHom.Locally import Mathlib.RingTheory.RingHom.StandardSmooth import Mathlib.RingTheory.RingHom.Surjective +import Mathlib.RingTheory.RingHom.Unramified import Mathlib.RingTheory.RingHomProperties import Mathlib.RingTheory.RingInvo import Mathlib.RingTheory.RootsOfUnity.AlgebraicallyClosed @@ -4657,6 +4669,7 @@ import Mathlib.RingTheory.SimpleRing.Field import Mathlib.RingTheory.SimpleRing.Matrix import Mathlib.RingTheory.Smooth.Basic import Mathlib.RingTheory.Smooth.Kaehler +import Mathlib.RingTheory.Smooth.Local import Mathlib.RingTheory.Smooth.Pi import Mathlib.RingTheory.Smooth.StandardSmooth import Mathlib.RingTheory.Support @@ -4690,6 +4703,7 @@ import Mathlib.RingTheory.UniqueFactorizationDomain.NormalizedFactors import Mathlib.RingTheory.Unramified.Basic import Mathlib.RingTheory.Unramified.Field import Mathlib.RingTheory.Unramified.Finite +import Mathlib.RingTheory.Unramified.Locus import Mathlib.RingTheory.Unramified.Pi import Mathlib.RingTheory.Valuation.AlgebraInstances import Mathlib.RingTheory.Valuation.Basic @@ -5396,6 +5410,7 @@ import Mathlib.Topology.Order.Hom.Basic import Mathlib.Topology.Order.Hom.Esakia import Mathlib.Topology.Order.IntermediateValue import Mathlib.Topology.Order.IsLUB +import Mathlib.Topology.Order.IsLocallyClosed import Mathlib.Topology.Order.Lattice import Mathlib.Topology.Order.LawsonTopology import Mathlib.Topology.Order.LeftRight diff --git a/Mathlib/Algebra/AddConstMap/Basic.lean b/Mathlib/Algebra/AddConstMap/Basic.lean index 25685a81e92e9..f71b10f9c82ba 100644 --- a/Mathlib/Algebra/AddConstMap/Basic.lean +++ b/Mathlib/Algebra/AddConstMap/Basic.lean @@ -87,7 +87,7 @@ theorem map_add_one [AddMonoidWithOne G] [Add H] [AddConstMapClass F G H 1 b] @[scoped simp] theorem map_add_ofNat' [AddMonoidWithOne G] [AddMonoid H] [AddConstMapClass F G H 1 b] (f : F) (x : G) (n : ℕ) [n.AtLeastTwo] : - f (x + no_index (OfNat.ofNat n)) = f x + (OfNat.ofNat n : ℕ) • b := + f (x + ofNat(n)) = f x + (OfNat.ofNat n : ℕ) • b := map_add_nat' f x n theorem map_add_nat [AddMonoidWithOne G] [AddMonoidWithOne H] [AddConstMapClass F G H 1 1] @@ -181,7 +181,7 @@ theorem map_sub_nat' [AddGroupWithOne G] [AddGroup H] [AddConstMapClass F G H 1 @[scoped simp] theorem map_sub_ofNat' [AddGroupWithOne G] [AddGroup H] [AddConstMapClass F G H 1 b] (f : F) (x : G) (n : ℕ) [n.AtLeastTwo] : - f (x - no_index (OfNat.ofNat n)) = f x - OfNat.ofNat n • b := + f (x - ofNat(n)) = f x - OfNat.ofNat n • b := map_sub_nat' f x n @[scoped simp] diff --git a/Mathlib/Algebra/Algebra/NonUnitalHom.lean b/Mathlib/Algebra/Algebra/NonUnitalHom.lean index 3ab73048ff9a8..88d007dcb7961 100644 --- a/Mathlib/Algebra/Algebra/NonUnitalHom.lean +++ b/Mathlib/Algebra/Algebra/NonUnitalHom.lean @@ -474,7 +474,7 @@ lemma coe_restrictScalars' (f : A →ₙₐ[S] B) : (f.restrictScalars R : A → theorem restrictScalars_injective : Function.Injective (restrictScalars R : (A →ₙₐ[S] B) → A →ₙₐ[R] B) := - fun _ _ h ↦ ext (congr_fun h : _) + fun _ _ h ↦ ext (congr_fun h :) end NonUnitalAlgHom diff --git a/Mathlib/Algebra/Algebra/NonUnitalSubalgebra.lean b/Mathlib/Algebra/Algebra/NonUnitalSubalgebra.lean index bea24677990a7..f4ac70dcbba21 100644 --- a/Mathlib/Algebra/Algebra/NonUnitalSubalgebra.lean +++ b/Mathlib/Algebra/Algebra/NonUnitalSubalgebra.lean @@ -467,7 +467,7 @@ theorem coe_codRestrict (f : F) (S : NonUnitalSubalgebra R B) (hf : ∀ x, f x theorem injective_codRestrict (f : F) (S : NonUnitalSubalgebra R B) (hf : ∀ x : A, f x ∈ S) : Function.Injective (NonUnitalAlgHom.codRestrict f S hf) ↔ Function.Injective f := - ⟨fun H _x _y hxy => H <| Subtype.eq hxy, fun H _x _y hxy => H (congr_arg Subtype.val hxy : _)⟩ + ⟨fun H _x _y hxy => H <| Subtype.eq hxy, fun H _x _y hxy => H (congr_arg Subtype.val hxy :)⟩ /-- Restrict the codomain of an `NonUnitalAlgHom` `f` to `f.range`. diff --git a/Mathlib/Algebra/Algebra/Operations.lean b/Mathlib/Algebra/Algebra/Operations.lean index 01b7c2e0ecb3f..7a693780a741f 100644 --- a/Mathlib/Algebra/Algebra/Operations.lean +++ b/Mathlib/Algebra/Algebra/Operations.lean @@ -144,10 +144,6 @@ theorem smul_mono_left (h : I ≤ J) : I • N ≤ J • N := instance : CovariantClass (Submodule R A) (Submodule R M) HSMul.hSMul LE.le := ⟨fun _ _ => smul_mono le_rfl⟩ -@[deprecated smul_mono_right (since := "2024-03-31")] -protected theorem smul_mono_right (h : N ≤ P) : I • N ≤ I • P := - _root_.smul_mono_right I h - variable (I J N P) @[simp] @@ -180,10 +176,6 @@ protected theorem smul_assoc {B} [Semiring B] [Module R B] [Module A B] [Module (fun j hj n hn ↦ (smul_assoc r j n).symm ▸ smul_mem_smul (smul_mem_smul hr hj) hn) fun m₁ m₂ ↦ (smul_add r m₁ m₂) ▸ add_mem) -@[deprecated smul_inf_le (since := "2024-03-31")] -protected theorem smul_inf_le (M₁ M₂ : Submodule R M) : - I • (M₁ ⊓ M₂) ≤ I • M₁ ⊓ I • M₂ := smul_inf_le _ _ _ - theorem smul_iSup {ι : Sort*} {I : Submodule R A} {t : ι → Submodule R M} : I • (⨆ i, t i)= ⨆ i, I • t i := toAddSubmonoid_injective <| by @@ -196,11 +188,6 @@ theorem iSup_smul {ι : Sort*} {t : ι → Submodule R A} {N : Submodule R M} : (by simp_rw [zero_smul]; apply zero_mem) fun x y ↦ by simp_rw [add_smul]; apply add_mem) (iSup_le fun i ↦ Submodule.smul_mono_left <| le_iSup _ i) -@[deprecated smul_iInf_le (since := "2024-03-31")] -protected theorem smul_iInf_le {ι : Sort*} {I : Submodule R A} {t : ι → Submodule R M} : - I • iInf t ≤ ⨅ i, I • t i := - smul_iInf_le - protected theorem one_smul : (1 : Submodule R A) • N = N := by refine le_antisymm (smul_le.mpr fun r hr m hm ↦ ?_) fun m hm ↦ ?_ · obtain ⟨r, rfl⟩ := hr diff --git a/Mathlib/Algebra/Algebra/RestrictScalars.lean b/Mathlib/Algebra/Algebra/RestrictScalars.lean index ba1e455fa5482..20a38263fd74c 100644 --- a/Mathlib/Algebra/Algebra/RestrictScalars.lean +++ b/Mathlib/Algebra/Algebra/RestrictScalars.lean @@ -124,7 +124,7 @@ instance RestrictScalars.opModule [Module Sᵐᵒᵖ M] : Module Rᵐᵒᵖ (Res instance RestrictScalars.isCentralScalar [Module S M] [Module Sᵐᵒᵖ M] [IsCentralScalar S M] : IsCentralScalar R (RestrictScalars R S M) where - op_smul_eq_smul r _x := (op_smul_eq_smul (algebraMap R S r) (_ : M) : _) + op_smul_eq_smul r _x := (op_smul_eq_smul (algebraMap R S r) (_ : M) :) /-- The `R`-algebra homomorphism from the original coefficient algebra `S` to endomorphisms of `RestrictScalars R S M`. diff --git a/Mathlib/Algebra/Algebra/Subalgebra/Basic.lean b/Mathlib/Algebra/Algebra/Subalgebra/Basic.lean index 20f8a381665c2..edf843664e624 100644 --- a/Mathlib/Algebra/Algebra/Subalgebra/Basic.lean +++ b/Mathlib/Algebra/Algebra/Subalgebra/Basic.lean @@ -520,7 +520,7 @@ theorem coe_codRestrict (f : A →ₐ[R] B) (S : Subalgebra R B) (hf : ∀ x, f theorem injective_codRestrict (f : A →ₐ[R] B) (S : Subalgebra R B) (hf : ∀ x, f x ∈ S) : Function.Injective (f.codRestrict S hf) ↔ Function.Injective f := - ⟨fun H _x _y hxy => H <| Subtype.eq hxy, fun H _x _y hxy => H (congr_arg Subtype.val hxy : _)⟩ + ⟨fun H _x _y hxy => H <| Subtype.eq hxy, fun H _x _y hxy => H (congr_arg Subtype.val hxy :)⟩ /-- Restrict the codomain of an `AlgHom` `f` to `f.range`. @@ -819,7 +819,7 @@ noncomputable def botEquivOfInjective (h : Function.Injective (algebraMap R A)) (⊥ : Subalgebra R A) ≃ₐ[R] R := AlgEquiv.symm <| AlgEquiv.ofBijective (Algebra.ofId R _) - ⟨fun _x _y hxy => h (congr_arg Subtype.val hxy : _), fun ⟨_y, x, hx⟩ => ⟨x, Subtype.eq hx⟩⟩ + ⟨fun _x _y hxy => h (congr_arg Subtype.val hxy :), fun ⟨_y, x, hx⟩ => ⟨x, Subtype.eq hx⟩⟩ /-- The bottom subalgebra is isomorphic to the field. -/ @[simps! symm_apply] @@ -995,7 +995,7 @@ instance isScalarTower_left [SMul α β] [SMul A α] [SMul A β] [IsScalarTower instance isScalarTower_mid {R S T : Type*} [CommSemiring R] [Semiring S] [AddCommMonoid T] [Algebra R S] [Module R T] [Module S T] [IsScalarTower R S T] (S' : Subalgebra R S) : IsScalarTower R S' T := - ⟨fun _x y _z => (smul_assoc _ (y : S) _ : _)⟩ + ⟨fun _x y _z => smul_assoc _ (y : S) _⟩ instance [SMul A α] [FaithfulSMul A α] (S : Subalgebra R A) : FaithfulSMul S α := inferInstanceAs (FaithfulSMul S.toSubsemiring α) diff --git a/Mathlib/Algebra/Algebra/Subalgebra/Tower.lean b/Mathlib/Algebra/Algebra/Subalgebra/Tower.lean index c34c387d80077..ebe721d031b4f 100644 --- a/Mathlib/Algebra/Algebra/Subalgebra/Tower.lean +++ b/Mathlib/Algebra/Algebra/Subalgebra/Tower.lean @@ -58,7 +58,7 @@ variable [Algebra R A] [IsScalarTower R S A] instance subalgebra' (S₀ : Subalgebra R S) : IsScalarTower R S₀ A := @IsScalarTower.of_algebraMap_eq R S₀ A _ _ _ _ _ _ fun _ ↦ - (IsScalarTower.algebraMap_apply R S A _ : _) + (IsScalarTower.algebraMap_apply R S A _ :) end Semiring diff --git a/Mathlib/Algebra/Algebra/Tower.lean b/Mathlib/Algebra/Algebra/Tower.lean index ba6b199b4aad6..16188b3a49e9b 100644 --- a/Mathlib/Algebra/Algebra/Tower.lean +++ b/Mathlib/Algebra/Algebra/Tower.lean @@ -200,7 +200,7 @@ theorem coe_restrictScalars' (f : A →ₐ[S] B) : (restrictScalars R f : A → theorem restrictScalars_injective : Function.Injective (restrictScalars R : (A →ₐ[S] B) → A →ₐ[R] B) := fun _ _ h => - AlgHom.ext (AlgHom.congr_fun h : _) + AlgHom.ext (AlgHom.congr_fun h :) end AlgHom @@ -223,7 +223,7 @@ theorem coe_restrictScalars' (f : A ≃ₐ[S] B) : (restrictScalars R f : A → theorem restrictScalars_injective : Function.Injective (restrictScalars R : (A ≃ₐ[S] B) → A ≃ₐ[R] B) := fun _ _ h => - AlgEquiv.ext (AlgEquiv.congr_fun h : _) + AlgEquiv.ext (AlgEquiv.congr_fun h :) end AlgEquiv diff --git a/Mathlib/Algebra/Associated/Basic.lean b/Mathlib/Algebra/Associated/Basic.lean index 4484964fb4604..12ae9d5fdd160 100644 --- a/Mathlib/Algebra/Associated/Basic.lean +++ b/Mathlib/Algebra/Associated/Basic.lean @@ -581,7 +581,6 @@ theorem mk_le_mk_of_dvd {a b : M} : a ∣ b → Associates.mk a ≤ Associates.m theorem mk_le_mk_iff_dvd {a b : M} : Associates.mk a ≤ Associates.mk b ↔ a ∣ b := mk_dvd_mk -@[deprecated (since := "2024-03-16")] alias mk_le_mk_iff_dvd_iff := mk_le_mk_iff_dvd @[simp] theorem isPrimal_mk {a : M} : IsPrimal (Associates.mk a) ↔ IsPrimal a := by @@ -591,7 +590,6 @@ theorem isPrimal_mk {a : M} : IsPrimal (Associates.mk a) ↔ IsPrimal a := by exact ⟨a₁, a₂ * u, h₁, Units.mul_right_dvd.mpr h₂, mul_assoc _ _ _⟩ · exact ⟨a₁, a₂, h₁, h₂, congr_arg _ eq⟩ -@[deprecated (since := "2024-03-16")] alias isPrimal_iff := isPrimal_mk @[simp] theorem decompositionMonoid_iff : DecompositionMonoid (Associates M) ↔ DecompositionMonoid M := by diff --git a/Mathlib/Algebra/BigOperators/Group/Multiset.lean b/Mathlib/Algebra/BigOperators/Group/Multiset.lean index a93afea870a07..9f83c37dae0a1 100644 --- a/Mathlib/Algebra/BigOperators/Group/Multiset.lean +++ b/Mathlib/Algebra/BigOperators/Group/Multiset.lean @@ -5,7 +5,8 @@ Authors: Mario Carneiro -/ import Mathlib.Algebra.BigOperators.Group.List import Mathlib.Algebra.Group.Prod -import Mathlib.Data.Multiset.Basic +import Mathlib.Algebra.Order.Group.Multiset +import Mathlib.Algebra.Order.Sub.Unbundled.Basic /-! # Sums and products over multisets diff --git a/Mathlib/Algebra/BigOperators/GroupWithZero/Action.lean b/Mathlib/Algebra/BigOperators/GroupWithZero/Action.lean index 6f377d4ff77be..af5e65fe47cb5 100644 --- a/Mathlib/Algebra/BigOperators/GroupWithZero/Action.lean +++ b/Mathlib/Algebra/BigOperators/GroupWithZero/Action.lean @@ -3,11 +3,11 @@ Copyright (c) 2020 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ +import Mathlib.Algebra.BigOperators.Finprod import Mathlib.Algebra.BigOperators.Group.Finset import Mathlib.Algebra.GroupWithZero.Action.Defs +import Mathlib.Algebra.Order.Group.Multiset import Mathlib.Data.Finset.Basic -import Mathlib.Data.Multiset.Basic -import Mathlib.Algebra.BigOperators.Finprod /-! # Lemmas about group actions on big operators diff --git a/Mathlib/Algebra/Category/Grp/Colimits.lean b/Mathlib/Algebra/Category/Grp/Colimits.lean index 1164b843e4f0b..a7d7824126397 100644 --- a/Mathlib/Algebra/Category/Grp/Colimits.lean +++ b/Mathlib/Algebra/Category/Grp/Colimits.lean @@ -1,37 +1,29 @@ /- Copyright (c) 2019 Kim Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Kim Morrison +Authors: Kim Morrison, Sophie Morel -/ import Mathlib.Algebra.Category.Grp.Preadditive -import Mathlib.CategoryTheory.Limits.Shapes.Kernels -import Mathlib.CategoryTheory.Limits.Shapes.FiniteLimits +import Mathlib.Algebra.Equiv.TransferInstance import Mathlib.CategoryTheory.ConcreteCategory.Elementwise +import Mathlib.Data.DFinsupp.BigOperators +import Mathlib.Data.DFinsupp.Small import Mathlib.GroupTheory.QuotientGroup.Defs - /-! # The category of additive commutative groups has all colimits. -This file uses a "pre-automated" approach, just as for `Algebra.Category.MonCat.Colimits`. -It is a very uniform approach, that conceivably could be synthesised directly -by a tactic that analyses the shape of `AddCommGroup` and `MonoidHom`. +This file constructs colimits in the categpry of additive commutative groups, as +quotients of finitely supported functions. -TODO: -In fact, in `AddCommGrp` there is a much nicer model of colimits as quotients -of finitely supported functions, and we really should implement this as well (or instead). -/ - universe w u v open CategoryTheory Limits --- [ROBOT VOICE]: --- You should pretend for now that this file was automatically generated. --- It follows the same template as colimits in Mon. namespace AddCommGrp -variable {J : Type u} [Category.{v} J] (F : J ⥤ AddCommGrp.{max u v w}) +variable {J : Type u} [Category.{v} J] (F : J ⥤ AddCommGrp.{w}) namespace Colimits @@ -42,220 +34,168 @@ then taking the quotient by the abelian group laws within each abelian group, and the identifications given by the morphisms in the diagram. -/ -/-- An inductive type representing all group expressions (without relations) -on a collection of types indexed by the objects of `J`. --/ -inductive Prequotient - -- There's always `of` - | of : ∀ (j : J) (_ : F.obj j), Prequotient - -- Then one generator for each operation - | zero : Prequotient - | neg : Prequotient → Prequotient - | add : Prequotient → Prequotient → Prequotient - -instance : Inhabited (Prequotient.{w} F) := - ⟨Prequotient.zero⟩ - -open Prequotient - -/-- The relation on `Prequotient` saying when two expressions are equal -because of the abelian group laws, or -because one element is mapped to another by a morphism in the diagram. +/-- +The relations between elements of the direct sum of the `F.obj j` given by the +morphisms in the diagram `J`. -/ -inductive Relation : Prequotient.{w} F → Prequotient.{w} F → Prop - -- Make it an equivalence relation: - | refl : ∀ x, Relation x x - | symm : ∀ (x y) (_ : Relation x y), Relation y x - | trans : ∀ (x y z) (_ : Relation x y) (_ : Relation y z), Relation x z - -- There's always a `map` relation - | map : ∀ (j j' : J) (f : j ⟶ j') (x : F.obj j), Relation (Prequotient.of j' (F.map f x)) - (Prequotient.of j x) - -- Then one relation per operation, describing the interaction with `of` - | zero : ∀ j, Relation (Prequotient.of j 0) zero - | neg : ∀ (j) (x : F.obj j), Relation (Prequotient.of j (-x)) (neg (Prequotient.of j x)) - | add : ∀ (j) (x y : F.obj j), Relation (Prequotient.of j (x + y)) (add (Prequotient.of j x) - (Prequotient.of j y)) - -- Then one relation per argument of each operation - | neg_1 : ∀ (x x') (_ : Relation x x'), Relation (neg x) (neg x') - | add_1 : ∀ (x x' y) (_ : Relation x x'), Relation (add x y) (add x' y) - | add_2 : ∀ (x y y') (_ : Relation y y'), Relation (add x y) (add x y') - -- And one relation per axiom - | zero_add : ∀ x, Relation (add zero x) x - | add_zero : ∀ x, Relation (add x zero) x - | neg_add_cancel : ∀ x, Relation (add (neg x) x) zero - | add_comm : ∀ x y, Relation (add x y) (add y x) - | add_assoc : ∀ x y z, Relation (add (add x y) z) (add x (add y z)) +abbrev Relations [DecidableEq J] : AddSubgroup (DFinsupp (fun j ↦ F.obj j)) := + AddSubgroup.closure {x | ∃ (j j' : J) (u : j ⟶ j') (a : F.obj j), + x = DFinsupp.single j' (F.map u a) - DFinsupp.single j a} /-- -The setoid corresponding to group expressions modulo abelian group relations and identifications. +The candidate for the colimit of `F`, defined as the quotient of the direct sum +of the commutative groups `F.obj j` by the relations given by the morphisms in +the diagram. -/ -def colimitSetoid : Setoid (Prequotient.{w} F) where - r := Relation F - iseqv := ⟨Relation.refl, fun r => Relation.symm _ _ r, fun r => Relation.trans _ _ _ r⟩ +def Quot [DecidableEq J] : Type (max u w) := + DFinsupp (fun j ↦ F.obj j) ⧸ Relations F -attribute [instance] colimitSetoid +instance [DecidableEq J] : AddCommGroup (Quot F) := + QuotientAddGroup.Quotient.addCommGroup (Relations F) -/-- The underlying type of the colimit of a diagram in `AddCommGrp`. +/-- Inclusion of `F.obj j` into the candidate colimit. -/ -def ColimitType : Type max u v w := - Quotient (colimitSetoid.{w} F) - -instance : Zero (ColimitType.{w} F) where - zero := Quotient.mk _ zero - -instance : Neg (ColimitType.{w} F) where - neg := Quotient.map neg Relation.neg_1 - -instance : Add (ColimitType.{w} F) where - add := Quotient.map₂ add <| fun _x x' rx y _y' ry => - Setoid.trans (Relation.add_1 _ _ y rx) (Relation.add_2 x' _ _ ry) - -instance : AddCommGroup (ColimitType.{w} F) where - zero_add := Quotient.ind <| fun _ => Quotient.sound <| Relation.zero_add _ - add_zero := Quotient.ind <| fun _ => Quotient.sound <| Relation.add_zero _ - neg_add_cancel := Quotient.ind <| fun _ => Quotient.sound <| Relation.neg_add_cancel _ - add_comm := Quotient.ind₂ <| fun _ _ => Quotient.sound <| Relation.add_comm _ _ - add_assoc := Quotient.ind <| fun _ => Quotient.ind₂ <| fun _ _ => - Quotient.sound <| Relation.add_assoc _ _ _ - nsmul := nsmulRec - zsmul := zsmulRec - -instance ColimitTypeInhabited : Inhabited (ColimitType.{w} F) := ⟨0⟩ - -@[simp] -theorem quot_zero : Quot.mk Setoid.r zero = (0 : ColimitType.{w} F) := - rfl +def Quot.ι [DecidableEq J] (j : J) : F.obj j →+ Quot F := + (QuotientAddGroup.mk' _).comp (DFinsupp.singleAddHom (fun j ↦ F.obj j) j) + +lemma Quot.addMonoidHom_ext [DecidableEq J] {α : Type*} [AddMonoid α] {f g : Quot F →+ α} + (h : ∀ (j : J) (x : F.obj j), f (Quot.ι F j x) = g (Quot.ι F j x)) : f = g := + QuotientAddGroup.addMonoidHom_ext _ (DFinsupp.addHom_ext h) + +variable (c : Cocone F) + +/-- (implementation detail) Part of the universal property of the colimit cocone, but without + assuming that `Quot F` lives in the correct universe. -/ +def Quot.desc [DecidableEq J] : Quot.{w} F →+ c.pt := by + refine QuotientAddGroup.lift _ (DFinsupp.sumAddHom c.ι.app) ?_ + dsimp + rw [AddSubgroup.closure_le] + intro _ ⟨_, _, _, _, eq⟩ + rw [eq] + simp only [SetLike.mem_coe, AddMonoidHom.mem_ker, map_sub, DFinsupp.sumAddHom_single] + change (F.map _ ≫ c.ι.app _) _ - _ = 0 + rw [c.ι.naturality] + simp only [Functor.const_obj_obj, Functor.const_obj_map, Category.comp_id, sub_self] @[simp] -theorem quot_neg (x) : - -- Porting note: force Lean to treat `ColimitType F` no as `Quot _` - (by exact Quot.mk Setoid.r (neg x) : ColimitType.{w} F) = - -(by exact Quot.mk Setoid.r x) := - rfl +lemma Quot.ι_desc [DecidableEq J] (j : J) (x : F.obj j) : + Quot.desc F c (Quot.ι F j x) = c.ι.app j x := by + dsimp [desc, ι] + erw [QuotientAddGroup.lift_mk'] + simp @[simp] -theorem quot_add (x y) : - (by exact Quot.mk Setoid.r (add x y) : ColimitType.{w} F) = - -- Porting note: force Lean to treat `ColimitType F` no as `Quot _` - (by exact Quot.mk Setoid.r x) + (by exact Quot.mk Setoid.r y) := - rfl - -/-- The bundled abelian group giving the colimit of a diagram. -/ -def colimit : AddCommGrp := - AddCommGrp.of (ColimitType.{w} F) - -/-- The function from a given abelian group in the diagram to the colimit abelian group. -/ -def coconeFun (j : J) (x : F.obj j) : ColimitType.{w} F := - Quot.mk _ (Prequotient.of j x) - -/-- The group homomorphism from a given abelian group in the diagram to the colimit abelian -group. -/ -def coconeMorphism (j : J) : F.obj j ⟶ colimit.{w} F where - toFun := coconeFun F j - map_zero' := by apply Quot.sound; apply Relation.zero - map_add' := by intros; apply Quot.sound; apply Relation.add +lemma Quot.map_ι [DecidableEq J] {j j' : J} {f : j ⟶ j'} (x : F.obj j) : + Quot.ι F j' (F.map f x) = Quot.ι F j x := by + dsimp [ι] + refine eq_of_sub_eq_zero ?_ + erw [← (QuotientAddGroup.mk' (Relations F)).map_sub, ← AddMonoidHom.mem_ker] + rw [QuotientAddGroup.ker_mk'] + simp only [DFinsupp.singleAddHom_apply] + exact AddSubgroup.subset_closure ⟨j, j', f, x, rfl⟩ + +/-- (implementation detail) A morphism of commutative additive groups `Quot F →+ A` +induces a cocone on `F` as long as the universes work out. +-/ +@[simps] +def toCocone [DecidableEq J] {A : Type w} [AddCommGroup A] (f : Quot F →+ A) : Cocone F where + pt := AddCommGrp.of A + ι := { app := fun j => f.comp (Quot.ι F j) } + +lemma Quot.desc_toCocone_desc [DecidableEq J] {A : Type w} [AddCommGroup A] (f : Quot F →+ A) + (hc : IsColimit c) : (hc.desc (toCocone F f)).comp (Quot.desc F c) = f := by + refine Quot.addMonoidHom_ext F (fun j x ↦ ?_) + rw [AddMonoidHom.comp_apply, ι_desc] + change (c.ι.app j ≫ hc.desc (toCocone F f)) _ = _ + rw [hc.fac] + simp + +lemma Quot.desc_toCocone_desc_app [DecidableEq J] {A : Type w} [AddCommGroup A] (f : Quot F →+ A) + (hc : IsColimit c) (x : Quot F) : hc.desc (toCocone F f) (Quot.desc F c x) = f x := by + conv_rhs => rw [← Quot.desc_toCocone_desc F c f hc] + dsimp -@[simp] -theorem cocone_naturality {j j' : J} (f : j ⟶ j') : - F.map f ≫ coconeMorphism.{w} F j' = coconeMorphism F j := by - ext - apply Quot.sound - apply Relation.map +/-- +If `c` is a cocone of `F` such that `Quot.desc F c` is bijective, then `c` is a colimit +cocone of `F`. +-/ +noncomputable def isColimit_of_bijective_desc [DecidableEq J] + (h : Function.Bijective (Quot.desc F c)) : IsColimit c where + desc s := AddCommGrp.ofHom ((Quot.desc F s).comp (AddEquiv.ofBijective + (Quot.desc F c) h).symm.toAddMonoidHom) + fac s j := by + ext x + dsimp + conv_lhs => erw [← Quot.ι_desc F c j x] + rw [← AddEquiv.ofBijective_apply _ h, AddEquiv.symm_apply_apply] + simp only [Quot.ι_desc, Functor.const_obj_obj] + uniq s m hm := by + ext x + obtain ⟨x, rfl⟩ := h.2 x + dsimp + rw [← AddEquiv.ofBijective_apply _ h, AddEquiv.symm_apply_apply] + suffices eq : m.comp (AddEquiv.ofBijective (Quot.desc F c) h) = Quot.desc F s by + rw [← eq]; rfl + exact Quot.addMonoidHom_ext F (by simp [← hm]) + +/-- (internal implementation) The colimit cocone of a functor `F`, implemented as a quotient of +`DFinsupp (fun j ↦ F.obj j)`, under the assumption that said quotient is small. +-/ +@[simps] +noncomputable def colimitCocone [DecidableEq J] [Small.{w} (Quot.{w} F)] : Cocone F where + pt := AddCommGrp.of (Shrink (Quot F)) + ι := + { app j := + AddCommGrp.ofHom (Shrink.addEquiv.symm.toAddMonoidHom.comp (Quot.ι F j)) + naturality _ _ _ := by + ext + dsimp + simp only [Category.comp_id, ofHom_apply, AddMonoidHom.coe_comp, AddMonoidHom.coe_coe, + Function.comp_apply] + change Shrink.addEquiv.symm _ = _ + rw [Quot.map_ι] } @[simp] -theorem cocone_naturality_components (j j' : J) (f : j ⟶ j') (x : F.obj j) : - (coconeMorphism.{w} F j') (F.map f x) = (coconeMorphism F j) x := by - rw [← cocone_naturality F f] - rfl - -/-- The cocone over the proposed colimit abelian group. -/ -def colimitCocone : Cocone F where - pt := colimit.{w} F - ι := { app := coconeMorphism F } - -/-- The function from the free abelian group on the diagram to the cone point of any other -cocone. -/ -@[simp] -def descFunLift (s : Cocone F) : Prequotient.{w} F → s.pt - | Prequotient.of j x => (s.ι.app j) x - | zero => 0 - | neg x => -descFunLift s x - | add x y => descFunLift s x + descFunLift s y - -/-- The function from the colimit abelian group to the cone point of any other cocone. -/ -def descFun (s : Cocone F) : ColimitType.{w} F → s.pt := by - fapply Quot.lift - · exact descFunLift F s - · intro x y r - induction r with - | refl => rfl - | symm _ _ _ r_ih => exact r_ih.symm - | trans _ _ _ _ _ r_ih_h r_ih_k => exact Eq.trans r_ih_h r_ih_k - | map j j' f x => simpa only [descFunLift, Functor.const_obj_obj] using - DFunLike.congr_fun (s.ι.naturality f) x - | zero => simp - | neg => simp - | add => simp - | neg_1 _ _ _ r_ih => dsimp; rw [r_ih] - | add_1 _ _ _ _ r_ih => dsimp; rw [r_ih] - | add_2 _ _ _ _ r_ih => dsimp; rw [r_ih] - | zero_add => dsimp; rw [zero_add] - | add_zero => dsimp; rw [add_zero] - | neg_add_cancel => dsimp; rw [neg_add_cancel] - | add_comm => dsimp; rw [add_comm] - | add_assoc => dsimp; rw [add_assoc] - -/-- The group homomorphism from the colimit abelian group to the cone point of any other cocone. -/ -def descMorphism (s : Cocone F) : colimit.{w} F ⟶ s.pt where - toFun := descFun F s - map_zero' := rfl - map_add' x y := Quot.induction_on₂ x y fun _ _ ↦ rfl - -/-- Evidence that the proposed colimit is the colimit. -/ -def colimitCoconeIsColimit : IsColimit (colimitCocone.{w} F) where - desc s := descMorphism F s - uniq s m w := DFunLike.ext _ _ fun x => Quot.inductionOn x fun x => by - change (m : ColimitType F →+ s.pt) _ = (descMorphism F s : ColimitType F →+ s.pt) _ - induction x using Prequotient.recOn with - | of j x => exact DFunLike.congr_fun (w j) x - | zero => - dsimp only [quot_zero] - rw [map_zero, map_zero] - | neg x ih => - dsimp only [quot_neg] - rw [map_neg, map_neg, ih] - | add x y ihx ihy => - simp only [quot_add] - rw [m.map_add, (descMorphism F s).map_add, ihx, ihy] +theorem Quot.desc_colimitCocone [DecidableEq J] (F : J ⥤ AddCommGrp.{w}) [Small.{w} (Quot F)] : + Quot.desc F (colimitCocone F) = (Shrink.addEquiv (α := Quot F)).symm.toAddMonoidHom := by + refine Quot.addMonoidHom_ext F (fun j x ↦ ?_) + simp only [colimitCocone_pt, coe_of, AddEquiv.toAddMonoidHom_eq_coe, AddMonoidHom.coe_coe] + erw [Quot.ι_desc] + simp + +/-- (internal implementation) The fact that the candidate colimit cocone constructed in +`colimitCocone` is the colimit. +-/ +noncomputable def colimitCoconeIsColimit [DecidableEq J] [Small.{w} (Quot F)] : + IsColimit (colimitCocone F) := by + refine isColimit_of_bijective_desc F _ ?_ + rw [Quot.desc_colimitCocone] + exact Shrink.addEquiv.symm.bijective end Colimits -lemma hasColimit : HasColimit F := ⟨_, Colimits.colimitCoconeIsColimit.{w} F⟩ - -variable (J) - -lemma hasColimitsOfShape : HasColimitsOfShape J AddCommGrp.{max u v w} where - has_colimit F := hasColimit.{w} F +open Colimits -lemma hasColimitsOfSize : HasColimitsOfSize.{v, u} AddCommGrp.{max u v w} := - ⟨fun _ => hasColimitsOfShape.{w} _⟩ +lemma hasColimit_of_small_quot [DecidableEq J] (h : Small.{w} (Quot F)) : HasColimit F := + ⟨_, colimitCoconeIsColimit F⟩ -instance hasColimits : HasColimits AddCommGrp.{w} := hasColimitsOfSize.{w} +instance [DecidableEq J] [Small.{w} J] : Small.{w} (Quot F) := + small_of_surjective (QuotientAddGroup.mk'_surjective _) -instance : HasColimitsOfSize.{v, v} (AddCommGrpMax.{u, v}) := hasColimitsOfSize.{u} -instance : HasColimitsOfSize.{u, u} (AddCommGrpMax.{u, v}) := hasColimitsOfSize.{v} -instance : HasColimitsOfSize.{u, v} (AddCommGrpMax.{u, v}) := hasColimitsOfSize.{u} -instance : HasColimitsOfSize.{v, u} (AddCommGrpMax.{u, v}) := hasColimitsOfSize.{u} -instance : HasColimitsOfSize.{0, 0} (AddCommGrp.{u}) := hasColimitsOfSize.{u, 0, 0} +instance hasColimit [Small.{w} J] (F : J ⥤ AddCommGrp.{w}) : HasColimit F := by + classical + exact hasColimit_of_small_quot F inferInstance -example : HasColimits AddCommGrpMax.{v, u} := - inferInstance -example : HasColimits AddCommGrpMax.{u, v} := - inferInstance +/-- +If `J` is `w`-small, then any functor `J ⥤ AddCommGrp.{w}` has a colimit. +-/ +instance hasColimitsOfShape [Small.{w} J] : HasColimitsOfShape J (AddCommGrp.{w}) where -example : HasColimits AddCommGrp.{u} := - inferInstance +/-- The category of additive commutative groups has all small colimits. +-/ +instance (priority := 1300) hasColimitsOfSize [UnivLE.{u, w}] : + HasColimitsOfSize.{v, u} (AddCommGrp.{w}) where end AddCommGrp diff --git a/Mathlib/Algebra/Category/Grp/FilteredColimits.lean b/Mathlib/Algebra/Category/Grp/FilteredColimits.lean index a31573223addd..b68cb33f290ff 100644 --- a/Mathlib/Algebra/Category/Grp/FilteredColimits.lean +++ b/Mathlib/Algebra/Category/Grp/FilteredColimits.lean @@ -95,11 +95,9 @@ noncomputable instance colimitGroup : Group (G.{v, u} F) := inv_mul_cancel := fun x => by refine Quot.inductionOn x ?_; clear x; intro x obtain ⟨j, x⟩ := x - erw [colimit_inv_mk_eq, - colimit_mul_mk_eq (F ⋙ forget₂ Grp MonCat.{max v u}) ⟨j, _⟩ ⟨j, _⟩ j (𝟙 j) (𝟙 j), - colimit_one_eq (F ⋙ forget₂ Grp MonCat.{max v u}) j] - dsimp - erw [CategoryTheory.Functor.map_id, inv_mul_cancel] } + erw [colimit_inv_mk_eq] + erw [colimit_mul_mk_eq (F ⋙ forget₂ Grp MonCat.{max v u}) ⟨j, _⟩ ⟨j, _⟩ j (𝟙 j) (𝟙 j)] + simp [colimit_one_eq (F ⋙ forget₂ Grp MonCat.{max v u}) j] } /-- The bundled group giving the filtered colimit of a diagram. -/ @[to_additive "The bundled additive group giving the filtered colimit of a diagram."] diff --git a/Mathlib/Algebra/Category/Grp/ForgetCorepresentable.lean b/Mathlib/Algebra/Category/Grp/ForgetCorepresentable.lean index bf80cb3f9cbf7..5583ecdd7e331 100644 --- a/Mathlib/Algebra/Category/Grp/ForgetCorepresentable.lean +++ b/Mathlib/Algebra/Category/Grp/ForgetCorepresentable.lean @@ -4,9 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Joël Riou -/ import Mathlib.Algebra.Category.Grp.Basic -import Mathlib.Algebra.Group.ULift -import Mathlib.CategoryTheory.Yoneda -import Mathlib.Algebra.Category.MonCat.ForgetCorepresentable /-! # The forget functor is corepresentable diff --git a/Mathlib/Algebra/Category/Grp/Limits.lean b/Mathlib/Algebra/Category/Grp/Limits.lean index f5aac619ae4f9..af71b540c80c5 100644 --- a/Mathlib/Algebra/Category/Grp/Limits.lean +++ b/Mathlib/Algebra/Category/Grp/Limits.lean @@ -3,6 +3,8 @@ Copyright (c) 2020 Kim Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison -/ + +import Mathlib.Algebra.Category.MonCat.ForgetCorepresentable import Mathlib.Algebra.Category.Grp.ForgetCorepresentable import Mathlib.Algebra.Category.Grp.Preadditive import Mathlib.Algebra.Category.MonCat.Limits diff --git a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Monoidal.lean b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Monoidal.lean index 9547c9239b0fc..cd94fb824ae02 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Monoidal.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Monoidal.lean @@ -63,7 +63,7 @@ variable {M₁ M₂ M₃ M₄} @[simp] lemma tensorObj_map_tmul {X Y : Cᵒᵖ} (f : X ⟶ Y) (m₁ : M₁.obj X) (m₂ : M₂.obj X) : - DFunLike.coe (α := (M₁.obj X ⊗ M₂.obj X : _)) + DFunLike.coe (α := (M₁.obj X ⊗ M₂.obj X :)) (β := fun _ ↦ (ModuleCat.restrictScalars (R.map f).hom).obj (M₁.obj Y ⊗ M₂.obj Y)) ((tensorObj M₁ M₂).map f).hom (m₁ ⊗ₜ[R.obj X] m₂) = M₁.map f m₁ ⊗ₜ[R.obj Y] M₂.map f m₂ := rfl diff --git a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafify.lean b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafify.lean index d1dec86a2c040..440f20249af9e 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafify.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafify.lean @@ -318,7 +318,7 @@ noncomputable def sheafify : SheafOfModules.{v} R where def toSheafify : M₀ ⟶ (restrictScalars α).obj (sheafify α φ).val := homMk φ (fun X r₀ m₀ ↦ by simpa using (Sheafify.map_smul_eq α φ (α.app _ r₀) (φ.app _ m₀) (𝟙 _) - r₀ (by aesop) m₀ (by simp)).symm) + r₀ (by simp) m₀ (by simp)).symm) lemma toSheafify_app_apply (X : Cᵒᵖ) (x : M₀.obj X) : ((toSheafify α φ).app X).hom x = φ.app X x := rfl diff --git a/Mathlib/Algebra/Category/ModuleCat/Subobject.lean b/Mathlib/Algebra/Category/ModuleCat/Subobject.lean index 8b5e532a67127..5fb8394a9c465 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Subobject.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Subobject.lean @@ -85,8 +85,7 @@ theorem toKernelSubobject_arrow {M N : ModuleCat R} {f : M ⟶ N} (x : LinearMap suffices ((arrow ((kernelSubobject f))) ∘ (kernelSubobjectIso f ≪≫ kernelIsoKer f).inv) x = x by convert this rw [Iso.trans_inv, ← LinearMap.coe_comp, ← hom_comp, Category.assoc] - simp only [Category.assoc, kernelSubobject_arrow', kernelIsoKer_inv_kernel_ι] - aesop_cat + simp /-- An extensionality lemma showing that two elements of a cokernel by an image are equal if they differ by an element of the image. diff --git a/Mathlib/Algebra/Category/MonCat/ForgetCorepresentable.lean b/Mathlib/Algebra/Category/MonCat/ForgetCorepresentable.lean index 796d9ab458bb6..ce3e674738fab 100644 --- a/Mathlib/Algebra/Category/MonCat/ForgetCorepresentable.lean +++ b/Mathlib/Algebra/Category/MonCat/ForgetCorepresentable.lean @@ -21,16 +21,6 @@ open CategoryTheory Opposite namespace MonoidHom -/-- The equivalence `(β →* γ) ≃ (α →* γ)` obtained by precomposition with -a multiplicative equivalence `e : α ≃* β`. -/ -@[simps] -def precompEquiv {α β : Type*} [Monoid α] [Monoid β] (e : α ≃* β) (γ : Type*) [Monoid γ] : - (β →* γ) ≃ (α →* γ) where - toFun f := f.comp e - invFun g := g.comp e.symm - left_inv _ := by ext; simp - right_inv _ := by ext; simp - /-- The equivalence `(Multiplicative ℕ →* α) ≃ α` for any monoid `α`. -/ @[simps] def fromMultiplicativeNatEquiv (α : Type u) [Monoid α] : (Multiplicative ℕ →* α) ≃ α where @@ -49,16 +39,6 @@ end MonoidHom namespace AddMonoidHom -/-- The equivalence `(β →+ γ) ≃ (α →+ γ)` obtained by precomposition with -an additive equivalence `e : α ≃+ β`. -/ -@[simps] -def precompEquiv {α β : Type*} [AddMonoid α] [AddMonoid β] (e : α ≃+ β) (γ : Type*) [AddMonoid γ] : - (β →+ γ) ≃ (α →+ γ) where - toFun f := f.comp e - invFun g := g.comp e.symm - left_inv _ := by ext; simp - right_inv _ := by ext; simp - /-- The equivalence `(ℤ →+ α) ≃ α` for any additive group `α`. -/ @[simps] def fromNatEquiv (α : Type u) [AddMonoid α] : (ℕ →+ α) ≃ α where diff --git a/Mathlib/Algebra/Category/Ring/Constructions.lean b/Mathlib/Algebra/Category/Ring/Constructions.lean index 8bf3814986f64..19a99274584c3 100644 --- a/Mathlib/Algebra/Category/Ring/Constructions.lean +++ b/Mathlib/Algebra/Category/Ring/Constructions.lean @@ -283,7 +283,7 @@ def equalizerFork : Fork f g := def equalizerForkIsLimit : IsLimit (equalizerFork f g) := by fapply Fork.IsLimit.mk' intro s - use ofHom <| s.ι.hom.codRestrict _ fun x => (ConcreteCategory.congr_hom s.condition x : _) + use ofHom <| s.ι.hom.codRestrict _ fun x => (ConcreteCategory.congr_hom s.condition x :) constructor · ext rfl @@ -327,7 +327,7 @@ instance equalizer_ι_is_local_ring_hom' (F : WalkingParallelPairᵒᵖ ⥤ Comm have : _ = limit.π F (walkingParallelPairOpEquiv.functor.obj _) := (limit.isoLimitCone_inv_π ⟨_, IsLimit.whiskerEquivalence (limit.isLimit F) walkingParallelPairOpEquiv⟩ - WalkingParallelPair.zero : _) + WalkingParallelPair.zero :) erw [← this] -- note: this was not needed before https://github.com/leanprover-community/mathlib4/pull/19757 haveI : IsLocalHom (limit.π (walkingParallelPairOpEquiv.functor ⋙ F) zero).hom := by @@ -370,8 +370,8 @@ def pullbackConeIsLimit {A B C : CommRingCat.{u}} (f : A ⟶ C) (g : B ⟶ C) : · intro s m e₁ e₂ refine hom_ext <| RingHom.ext fun (x : s.pt) => Subtype.ext ?_ change (m x).1 = (_, _) - have eq1 := (congr_arg (fun f : s.pt →+* A => f x) (congrArg Hom.hom e₁) : _) - have eq2 := (congr_arg (fun f : s.pt →+* B => f x) (congrArg Hom.hom e₂) : _) + have eq1 := (congr_arg (fun f : s.pt →+* A => f x) (congrArg Hom.hom e₁) :) + have eq2 := (congr_arg (fun f : s.pt →+* B => f x) (congrArg Hom.hom e₂) :) rw [← eq1, ← eq2] rfl diff --git a/Mathlib/Algebra/Category/Ring/Instances.lean b/Mathlib/Algebra/Category/Ring/Instances.lean index 4fcf3b005de52..cc50fa4a7692e 100644 --- a/Mathlib/Algebra/Category/Ring/Instances.lean +++ b/Mathlib/Algebra/Category/Ring/Instances.lean @@ -34,7 +34,7 @@ instance Localization.epi {R : Type*} [CommRing R] (M : Submonoid R) : IsLocalization.epi M _ instance Localization.epi' {R : CommRingCat} (M : Submonoid R) : - @Epi CommRingCat _ R _ (CommRingCat.ofHom <| algebraMap R <| Localization M : _) := by + @Epi CommRingCat _ R _ (CommRingCat.ofHom <| algebraMap R <| Localization M :) := by rcases R with ⟨α, str⟩ exact IsLocalization.epi M _ diff --git a/Mathlib/Algebra/Category/Semigrp/Basic.lean b/Mathlib/Algebra/Category/Semigrp/Basic.lean index d30f2b19f232e..e6e54e5ccab88 100644 --- a/Mathlib/Algebra/Category/Semigrp/Basic.lean +++ b/Mathlib/Algebra/Category/Semigrp/Basic.lean @@ -44,7 +44,7 @@ namespace MagmaCat @[to_additive] instance bundledHom : BundledHom @MulHom := ⟨@MulHom.toFun, @MulHom.id, @MulHom.comp, - by intros; apply @DFunLike.coe_injective, by aesop_cat, by aesop_cat⟩ + by intros; apply @DFunLike.coe_injective, by aesop_cat, by simp⟩ -- Porting note: deriving failed for `ConcreteCategory`, -- "default handlers have not been implemented yet" diff --git a/Mathlib/Algebra/CharP/Reduced.lean b/Mathlib/Algebra/CharP/Reduced.lean index deea5f7b1eb53..90e5aa6b3aa17 100644 --- a/Mathlib/Algebra/CharP/Reduced.lean +++ b/Mathlib/Algebra/CharP/Reduced.lean @@ -44,6 +44,3 @@ theorem ExpChar.pow_prime_pow_mul_eq_one_iff (p k m : ℕ) [ExpChar R p] (x : R) rw [pow_mul'] convert ← (iterateFrobenius_inj R p k).eq_iff apply map_one - -@[deprecated (since := "2024-02-02")] -alias CharP.pow_prime_pow_mul_eq_one_iff := ExpChar.pow_prime_pow_mul_eq_one_iff diff --git a/Mathlib/Algebra/CharZero/Defs.lean b/Mathlib/Algebra/CharZero/Defs.lean index c82356e6fa6e7..f536a1bb1d7a4 100644 --- a/Mathlib/Algebra/CharZero/Defs.lean +++ b/Mathlib/Algebra/CharZero/Defs.lean @@ -89,20 +89,20 @@ namespace OfNat variable [AddMonoidWithOne R] [CharZero R] -@[simp] lemma ofNat_ne_zero (n : ℕ) [n.AtLeastTwo] : (no_index (ofNat n) : R) ≠ 0 := +@[simp] lemma ofNat_ne_zero (n : ℕ) [n.AtLeastTwo] : (ofNat(n) : R) ≠ 0 := Nat.cast_ne_zero.2 (NeZero.ne n) -@[simp] lemma zero_ne_ofNat (n : ℕ) [n.AtLeastTwo] : 0 ≠ (no_index (ofNat n) : R) := +@[simp] lemma zero_ne_ofNat (n : ℕ) [n.AtLeastTwo] : 0 ≠ (ofNat(n) : R) := (ofNat_ne_zero n).symm -@[simp] lemma ofNat_ne_one (n : ℕ) [n.AtLeastTwo] : (no_index (ofNat n) : R) ≠ 1 := +@[simp] lemma ofNat_ne_one (n : ℕ) [n.AtLeastTwo] : (ofNat(n) : R) ≠ 1 := Nat.cast_ne_one.2 (Nat.AtLeastTwo.ne_one) -@[simp] lemma one_ne_ofNat (n : ℕ) [n.AtLeastTwo] : (1 : R) ≠ no_index (ofNat n) := +@[simp] lemma one_ne_ofNat (n : ℕ) [n.AtLeastTwo] : (1 : R) ≠ ofNat(n) := (ofNat_ne_one n).symm @[simp] lemma ofNat_eq_ofNat {m n : ℕ} [m.AtLeastTwo] [n.AtLeastTwo] : - (no_index (ofNat m) : R) = no_index (ofNat n) ↔ (ofNat m : ℕ) = ofNat n := + (ofNat(m) : R) = ofNat(n) ↔ (ofNat m : ℕ) = ofNat n := Nat.cast_inj end OfNat diff --git a/Mathlib/Algebra/DirectSum/Algebra.lean b/Mathlib/Algebra/DirectSum/Algebra.lean index 80c5c808cd07e..5db0ecc0efa46 100644 --- a/Mathlib/Algebra/DirectSum/Algebra.lean +++ b/Mathlib/Algebra/DirectSum/Algebra.lean @@ -129,7 +129,7 @@ theorem algHom_ext' ⦃f g : (⨁ i, A i) →ₐ[R] B⦄ theorem algHom_ext ⦃f g : (⨁ i, A i) →ₐ[R] B⦄ (h : ∀ i x, f (of A i x) = g (of A i x)) : f = g := algHom_ext' R A fun i => LinearMap.ext <| h i -/-- The piecewise multiplication from the `Mul` instance, as a bundled linear homomorphism. +/-- The piecewise multiplication from the `Mul` instance, as a bundled linear map. This is the graded version of `LinearMap.mul`, and the linear version of `DirectSum.gMulHom` -/ @[simps] diff --git a/Mathlib/Algebra/DirectSum/Internal.lean b/Mathlib/Algebra/DirectSum/Internal.lean index a75722faa88dc..fb7946b75ad16 100644 --- a/Mathlib/Algebra/DirectSum/Internal.lean +++ b/Mathlib/Algebra/DirectSum/Internal.lean @@ -337,7 +337,7 @@ instance instSemiring : Semiring (A 0) := (subsemiring A).toSemiring @[simp, norm_cast] theorem coe_natCast (n : ℕ) : (n : A 0) = (n : R) := rfl @[simp, norm_cast] theorem coe_ofNat (n : ℕ) [n.AtLeastTwo] : - (no_index (OfNat.ofNat n) : A 0) = (OfNat.ofNat n : R) := rfl + (ofNat(n) : A 0) = (OfNat.ofNat n : R) := rfl end Semiring diff --git a/Mathlib/Algebra/Field/Basic.lean b/Mathlib/Algebra/Field/Basic.lean index b022d2062981b..e98a7a85d09fc 100644 --- a/Mathlib/Algebra/Field/Basic.lean +++ b/Mathlib/Algebra/Field/Basic.lean @@ -73,47 +73,6 @@ protected theorem Commute.inv_add_inv (hab : Commute a b) (ha : a ≠ 0) (hb : b end DivisionSemiring -section DivisionMonoid - -variable [DivisionMonoid K] [HasDistribNeg K] {a b : K} - -theorem one_div_neg_one_eq_neg_one : (1 : K) / -1 = -1 := - have : -1 * -1 = (1 : K) := by rw [neg_mul_neg, one_mul] - Eq.symm (eq_one_div_of_mul_eq_one_right this) - -theorem one_div_neg_eq_neg_one_div (a : K) : 1 / -a = -(1 / a) := - calc - 1 / -a = 1 / (-1 * a) := by rw [neg_eq_neg_one_mul] - _ = 1 / a * (1 / -1) := by rw [one_div_mul_one_div_rev] - _ = 1 / a * -1 := by rw [one_div_neg_one_eq_neg_one] - _ = -(1 / a) := by rw [mul_neg, mul_one] - -theorem div_neg_eq_neg_div (a b : K) : b / -a = -(b / a) := - calc - b / -a = b * (1 / -a) := by rw [← inv_eq_one_div, division_def] - _ = b * -(1 / a) := by rw [one_div_neg_eq_neg_one_div] - _ = -(b * (1 / a)) := by rw [neg_mul_eq_mul_neg] - _ = -(b / a) := by rw [mul_one_div] - -theorem neg_div (a b : K) : -b / a = -(b / a) := by - rw [neg_eq_neg_one_mul, mul_div_assoc, ← neg_eq_neg_one_mul] - -@[field_simps] -theorem neg_div' (a b : K) : -(b / a) = -b / a := by simp [neg_div] - -@[simp] -theorem neg_div_neg_eq (a b : K) : -a / -b = a / b := by rw [div_neg_eq_neg_div, neg_div, neg_neg] - -theorem neg_inv : -a⁻¹ = (-a)⁻¹ := by rw [inv_eq_one_div, inv_eq_one_div, div_neg_eq_neg_div] - -theorem div_neg (a : K) : a / -b = -(a / b) := by rw [← div_neg_eq_neg_div] - -theorem inv_neg : (-a)⁻¹ = -a⁻¹ := by rw [neg_inv] - -theorem inv_neg_one : (-1 : K)⁻¹ = -1 := by rw [← neg_inv, inv_one] - -end DivisionMonoid - section DivisionRing variable [DivisionRing K] {a b c d : K} diff --git a/Mathlib/Algebra/GCDMonoid/Basic.lean b/Mathlib/Algebra/GCDMonoid/Basic.lean index 16298274aed73..57c84854496c4 100644 --- a/Mathlib/Algebra/GCDMonoid/Basic.lean +++ b/Mathlib/Algebra/GCDMonoid/Basic.lean @@ -805,9 +805,6 @@ theorem lcm_eq_of_associated_right [NormalizedGCDMonoid α] {m n : α} (h : Asso end LCM -@[deprecated (since := "2024-02-12")] alias GCDMonoid.prime_of_irreducible := Irreducible.prime -@[deprecated (since := "2024-02-12")] alias GCDMonoid.irreducible_iff_prime := irreducible_iff_prime - end GCDMonoid section UniqueUnit diff --git a/Mathlib/Algebra/GCDMonoid/IntegrallyClosed.lean b/Mathlib/Algebra/GCDMonoid/IntegrallyClosed.lean index c528c210ffca8..6a53eb7e53fbd 100644 --- a/Mathlib/Algebra/GCDMonoid/IntegrallyClosed.lean +++ b/Mathlib/Algebra/GCDMonoid/IntegrallyClosed.lean @@ -40,6 +40,6 @@ instance (priority := 100) GCDMonoid.toIsIntegrallyClosed exact (dvd_gcd this <| dvd_refl y).trans (gcd_pow_left_dvd_pow_gcd.trans <| pow_dvd_pow_of_dvd (isUnit_iff_dvd_one.1 hg) _) - use x * (this.unit⁻¹ : _) + use x * (this.unit⁻¹ :) erw [map_mul, ← Units.coe_map_inv, eq_comm, Units.eq_mul_inv_iff_mul_eq] exact he diff --git a/Mathlib/Algebra/GradedMonoid.lean b/Mathlib/Algebra/GradedMonoid.lean index a4aa2463b4229..85e9bb219662e 100644 --- a/Mathlib/Algebra/GradedMonoid.lean +++ b/Mathlib/Algebra/GradedMonoid.lean @@ -403,11 +403,11 @@ theorem List.dProd_nil (fι : α → ι) (fA : ∀ a, A (fι a)) : (List.nil : List α).dProd fι fA = GradedMonoid.GOne.one := rfl --- the `( : _)` in this lemma statement results in the type on the RHS not being unfolded, which +-- the `( :)` in this lemma statement results in the type on the RHS not being unfolded, which -- is nicer in the goal view. @[simp] theorem List.dProd_cons (fι : α → ι) (fA : ∀ a, A (fι a)) (a : α) (l : List α) : - (a :: l).dProd fι fA = (GradedMonoid.GMul.mul (fA a) (l.dProd fι fA) : _) := + (a :: l).dProd fι fA = (GradedMonoid.GMul.mul (fA a) (l.dProd fι fA) :) := rfl theorem GradedMonoid.mk_list_dProd (l : List α) (fι : α → ι) (fA : ∀ a, A (fι a)) : diff --git a/Mathlib/Algebra/Group/Action/Units.lean b/Mathlib/Algebra/Group/Action/Units.lean index f2a6ef61171a5..5c32fade58481 100644 --- a/Mathlib/Algebra/Group/Action/Units.lean +++ b/Mathlib/Algebra/Group/Action/Units.lean @@ -46,20 +46,20 @@ instance [Monoid M] [SMul M α] [FaithfulSMul M α] : FaithfulSMul Mˣ α where @[to_additive] instance instMulAction [Monoid M] [MulAction M α] : MulAction Mˣ α where - one_smul := (one_smul M : _) + one_smul := one_smul M mul_smul m n := mul_smul (m : M) n @[to_additive] instance smulCommClass_left [Monoid M] [SMul M α] [SMul N α] [SMulCommClass M N α] : - SMulCommClass Mˣ N α where smul_comm m n := (smul_comm (m : M) n : _) + SMulCommClass Mˣ N α where smul_comm m n := smul_comm (m : M) n @[to_additive] instance smulCommClass_right [Monoid N] [SMul M α] [SMul N α] [SMulCommClass M N α] : - SMulCommClass M Nˣ α where smul_comm m n := (smul_comm m (n : N) : _) + SMulCommClass M Nˣ α where smul_comm m n := smul_comm m (n : N) @[to_additive] instance [Monoid M] [SMul M N] [SMul M α] [SMul N α] [IsScalarTower M N α] : - IsScalarTower Mˣ N α where smul_assoc m n := (smul_assoc (m : M) n : _) + IsScalarTower Mˣ N α where smul_assoc m n := smul_assoc (m : M) n /-! ### Action of a group `G` on units of `M` -/ @@ -102,7 +102,7 @@ instance isScalarTower' [SMul G H] [Group G] [Group H] [Monoid M] [MulAction G M @[to_additive "Transfer `VAddAssocClass G M α` to `VAddAssocClass G (AddUnits M) α`."] instance isScalarTower'_left [Group G] [Monoid M] [MulAction G M] [SMul M α] [SMul G α] [SMulCommClass G M M] [IsScalarTower G M M] [IsScalarTower G M α] : - IsScalarTower G Mˣ α where smul_assoc g m := (smul_assoc g (m : M) : _) + IsScalarTower G Mˣ α where smul_assoc g m := smul_assoc g (m : M) -- Just to prove this transfers a particularly useful instance. example [Monoid M] [Monoid N] [MulAction M N] [SMulCommClass M N N] [IsScalarTower M N N] : diff --git a/Mathlib/Algebra/Group/Basic.lean b/Mathlib/Algebra/Group/Basic.lean index 1df564f491efd..bf22f85392085 100644 --- a/Mathlib/Algebra/Group/Basic.lean +++ b/Mathlib/Algebra/Group/Basic.lean @@ -797,14 +797,6 @@ theorem leftInverse_inv_mul_mul_right (c : G) : @[to_additive (attr := simp) natAbs_nsmul_eq_zero] lemma pow_natAbs_eq_one : a ^ n.natAbs = 1 ↔ a ^ n = 1 := by cases n <;> simp -set_option linter.existingAttributeWarning false in -@[to_additive, deprecated pow_natAbs_eq_one (since := "2024-02-14")] -lemma exists_pow_eq_one_of_zpow_eq_one (hn : n ≠ 0) (h : a ^ n = 1) : - ∃ n : ℕ, 0 < n ∧ a ^ n = 1 := ⟨_, Int.natAbs_pos.2 hn, pow_natAbs_eq_one.2 h⟩ - -attribute [deprecated natAbs_nsmul_eq_zero (since := "2024-02-14")] -exists_nsmul_eq_zero_of_zsmul_eq_zero - @[to_additive sub_nsmul] lemma pow_sub (a : G) {m n : ℕ} (h : n ≤ m) : a ^ (m - n) = a ^ m * (a ^ n)⁻¹ := eq_mul_inv_of_mul_eq <| by rw [← pow_add, Nat.sub_add_cancel h] @@ -1041,15 +1033,5 @@ theorem multiplicative_of_isTotal (p : α → Prop) (hswap : ∀ {a b}, p a → end multiplicative -@[deprecated (since := "2024-03-20")] alias div_mul_cancel' := div_mul_cancel -@[deprecated (since := "2024-03-20")] alias mul_div_cancel'' := mul_div_cancel_right -- The name `add_sub_cancel` was reused -- @[deprecated (since := "2024-03-20")] alias add_sub_cancel := add_sub_cancel_right -@[deprecated (since := "2024-03-20")] alias div_mul_cancel''' := div_mul_cancel_right -@[deprecated (since := "2024-03-20")] alias sub_add_cancel'' := sub_add_cancel_right -@[deprecated (since := "2024-03-20")] alias mul_div_cancel''' := mul_div_cancel_left -@[deprecated (since := "2024-03-20")] alias add_sub_cancel' := add_sub_cancel_left -@[deprecated (since := "2024-03-20")] alias mul_div_cancel'_right := mul_div_cancel -@[deprecated (since := "2024-03-20")] alias add_sub_cancel'_right := add_sub_cancel -@[deprecated (since := "2024-03-20")] alias div_mul_cancel'' := div_mul_cancel_left -@[deprecated (since := "2024-03-20")] alias sub_add_cancel' := sub_add_cancel_left diff --git a/Mathlib/Algebra/Group/Defs.lean b/Mathlib/Algebra/Group/Defs.lean index 324b405f14be4..9808a37d910f3 100644 --- a/Mathlib/Algebra/Group/Defs.lean +++ b/Mathlib/Algebra/Group/Defs.lean @@ -886,8 +886,6 @@ theorem zpow_natCast (a : G) : ∀ n : ℕ, a ^ (n : ℤ) = a ^ n _ = a ^ n * a := congrArg (· * a) (zpow_natCast a n) _ = a ^ (n + 1) := (pow_succ _ _).symm -@[deprecated (since := "2024-03-20")] alias zpow_coe_nat := zpow_natCast -@[deprecated (since := "2024-03-20")] alias coe_nat_zsmul := natCast_zsmul @[to_additive ofNat_zsmul] lemma zpow_ofNat (a : G) (n : ℕ) : a ^ (ofNat(n) : ℤ) = a ^ OfNat.ofNat n := diff --git a/Mathlib/Algebra/Group/Equiv/Basic.lean b/Mathlib/Algebra/Group/Equiv/Basic.lean index 4d516e138b15b..30605dc8e53f9 100644 --- a/Mathlib/Algebra/Group/Equiv/Basic.lean +++ b/Mathlib/Algebra/Group/Equiv/Basic.lean @@ -634,6 +634,34 @@ protected theorem map_div [Group G] [DivisionMonoid H] (h : G ≃* H) (x y : G) end MulEquiv +namespace MonoidHom + +/-- The equivalence `(β →* γ) ≃ (α →* γ)` obtained by precomposition with +a multiplicative equivalence `e : α ≃* β`. -/ +@[to_additive (attr := simps) +"The equivalence `(β →+ γ) ≃ (α →+ γ)` obtained by precomposition with +an additive equivalence `e : α ≃+ β`."] +def precompEquiv {α β : Type*} [Monoid α] [Monoid β] (e : α ≃* β) (γ : Type*) [Monoid γ] : + (β →* γ) ≃ (α →* γ) where + toFun f := f.comp e + invFun g := g.comp e.symm + left_inv _ := by ext; simp + right_inv _ := by ext; simp + +/-- The equivalence `(γ →* α) ≃ (γ →* β)` obtained by postcomposition with +a multiplicative equivalence `e : α ≃* β`. -/ +@[to_additive (attr := simps) +"The equivalence `(γ →+ α) ≃ (γ →+ β)` obtained by postcomposition with +an additive equivalence `e : α ≃+ β`."] +def postcompEquiv {α β : Type*} [Monoid α] [Monoid β] (e : α ≃* β) (γ : Type*) [Monoid γ] : + (γ →* α) ≃ (γ →* β) where + toFun f := e.toMonoidHom.comp f + invFun g := e.symm.toMonoidHom.comp g + left_inv _ := by ext; simp + right_inv _ := by ext; simp + +end MonoidHom + -- Porting note: we want to add -- `@[simps (config := .asFn)]` -- here, but it generates simp lemmas which aren't in simp normal form diff --git a/Mathlib/Algebra/Group/Opposite.lean b/Mathlib/Algebra/Group/Opposite.lean index 527a665d95d21..5d91ad7cacb87 100644 --- a/Mathlib/Algebra/Group/Opposite.lean +++ b/Mathlib/Algebra/Group/Opposite.lean @@ -556,6 +556,6 @@ theorem AddMonoidHom.mul_op_ext {α β} [AddZeroClass α] [AddZeroClass β] (f g f.comp (opAddEquiv : α ≃+ αᵐᵒᵖ).toAddMonoidHom = g.comp (opAddEquiv : α ≃+ αᵐᵒᵖ).toAddMonoidHom) : f = g := - AddMonoidHom.ext <| MulOpposite.rec' fun x => (DFunLike.congr_fun h : _) x + AddMonoidHom.ext <| MulOpposite.rec' fun x => (DFunLike.congr_fun h :) x end Ext diff --git a/Mathlib/Algebra/Group/Pi/Lemmas.lean b/Mathlib/Algebra/Group/Pi/Lemmas.lean index 1d2924faf728b..e9dd7d38a4ff0 100644 --- a/Mathlib/Algebra/Group/Pi/Lemmas.lean +++ b/Mathlib/Algebra/Group/Pi/Lemmas.lean @@ -69,7 +69,7 @@ def Pi.mulHom {γ : Type w} [∀ i, Mul (f i)] [Mul γ] (g : ∀ i, γ →ₙ* f theorem Pi.mulHom_injective {γ : Type w} [Nonempty I] [∀ i, Mul (f i)] [Mul γ] (g : ∀ i, γ →ₙ* f i) (hg : ∀ i, Function.Injective (g i)) : Function.Injective (Pi.mulHom g) := fun _ _ h => let ⟨i⟩ := ‹Nonempty I› - hg i ((funext_iff.mp h : _) i) + hg i ((funext_iff.mp h :) i) /-- A family of monoid homomorphisms `f a : γ →* β a` defines a monoid homomorphism `Pi.monoidHom f : γ →* Π a, β a` given by `Pi.monoidHom f x b = f b x`. -/ diff --git a/Mathlib/Algebra/Group/Prod.lean b/Mathlib/Algebra/Group/Prod.lean index 739752b4ea517..de72f571ab908 100644 --- a/Mathlib/Algebra/Group/Prod.lean +++ b/Mathlib/Algebra/Group/Prod.lean @@ -697,7 +697,7 @@ def embedProduct (α : Type*) [Monoid α] : αˣ →* α × αᵐᵒᵖ where @[to_additive] theorem embedProduct_injective (α : Type*) [Monoid α] : Function.Injective (embedProduct α) := - fun _ _ h => Units.ext <| (congr_arg Prod.fst h : _) + fun _ _ h => Units.ext <| (congr_arg Prod.fst h :) end Units diff --git a/Mathlib/Algebra/Group/Subgroup/Basic.lean b/Mathlib/Algebra/Group/Subgroup/Basic.lean index e3e7ad321eee8..954661ef44610 100644 --- a/Mathlib/Algebra/Group/Subgroup/Basic.lean +++ b/Mathlib/Algebra/Group/Subgroup/Basic.lean @@ -642,9 +642,7 @@ theorem map_equiv_normalizer_eq (H : Subgroup G) (f : G ≃* N) : simp only [mem_normalizer_iff, mem_map_equiv] rw [f.toEquiv.forall_congr] intro - erw [f.toEquiv.symm_apply_apply] - simp only [map_mul, map_inv] - erw [f.toEquiv.symm_apply_apply] + simp /-- The image of the normalizer is equal to the normalizer of the image of a bijective function. -/ diff --git a/Mathlib/Algebra/Group/Subgroup/Even.lean b/Mathlib/Algebra/Group/Subgroup/Even.lean index b4735654ebe05..488a9b76c3b59 100644 --- a/Mathlib/Algebra/Group/Subgroup/Even.lean +++ b/Mathlib/Algebra/Group/Subgroup/Even.lean @@ -13,73 +13,73 @@ This file defines the subgroup of squares / even elements in an abelian group. -/ namespace Subsemigroup -variable {S : Type*} [CommSemigroup S] {a : S} +variable {S : Type*} [CommSemigroup S] variable (S) in /-- -In a commutative semigroup `S`, `Subsemigroup.squareIn S` is the subsemigroup of squares in `S`. +In a commutative semigroup `S`, `Subsemigroup.square S` is the subsemigroup of squares in `S`. -/ @[to_additive -"In a commutative additive semigroup `S`, the type `AddSubsemigroup.evenIn S` -is the subsemigroup of even elements of `S`."] -def squareIn : Subsemigroup S where +"In a commutative additive semigroup `S`, `AddSubsemigroup.even S` +is the subsemigroup of even elements in `S`."] +def square : Subsemigroup S where carrier := {s : S | IsSquare s} mul_mem' := IsSquare.mul @[to_additive (attr := simp)] -theorem mem_squareIn : a ∈ squareIn S ↔ IsSquare a := Iff.rfl +theorem mem_square {a : S} : a ∈ square S ↔ IsSquare a := Iff.rfl @[to_additive (attr := simp, norm_cast)] -theorem coe_squareIn : squareIn S = {s : S | IsSquare s} := rfl +theorem coe_square : square S = {s : S | IsSquare s} := rfl end Subsemigroup namespace Submonoid -variable {M : Type*} [CommMonoid M] {a : M} +variable {M : Type*} [CommMonoid M] variable (M) in /-- -In a commutative monoid `M`, `Submonoid.squareIn M` is the submonoid of squares in `M`. +In a commutative monoid `M`, `Submonoid.square M` is the submonoid of squares in `M`. -/ @[to_additive -"In a commutative additive monoid `M`, the type `AddSubmonoid.evenIn M` -is the submonoid of even elements of `M`."] -def squareIn : Submonoid M where - __ := Subsemigroup.squareIn M +"In a commutative additive monoid `M`, `AddSubmonoid.even M` +is the submonoid of even elements in `M`."] +def square : Submonoid M where + __ := Subsemigroup.square M one_mem' := IsSquare.one @[to_additive (attr := simp)] -theorem squareIn_toSubsemigroup : (squareIn M).toSubsemigroup = .squareIn M := rfl +theorem square_toSubsemigroup : (square M).toSubsemigroup = .square M := rfl @[to_additive (attr := simp)] -theorem mem_squareIn : a ∈ squareIn M ↔ IsSquare a := Iff.rfl +theorem mem_square {a : M} : a ∈ square M ↔ IsSquare a := Iff.rfl @[to_additive (attr := simp, norm_cast)] -theorem coe_squareIn : squareIn M = {s : M | IsSquare s} := rfl +theorem coe_square : square M = {s : M | IsSquare s} := rfl end Submonoid namespace Subgroup -variable {G : Type*} [CommGroup G] {a : G} +variable {G : Type*} [CommGroup G] variable (G) in /-- -In an abelian group `G`, `Subgroup.squareIn G` is the subgroup of squares in `G`. +In an abelian group `G`, `Subgroup.square G` is the subgroup of squares in `G`. -/ @[to_additive -"In an abelian additive group `G`, the type `AddSubgroup.evenIn G` is +"In an abelian additive group `G`, `AddSubgroup.even G` is the subgroup of even elements in `G`."] -def squareIn : Subgroup G where - __ := Submonoid.squareIn G +def square : Subgroup G where + __ := Submonoid.square G inv_mem' := IsSquare.inv @[to_additive (attr := simp)] -theorem squareIn_toSubmonoid : (squareIn G).toSubmonoid = .squareIn G := rfl +theorem square_toSubmonoid : (square G).toSubmonoid = .square G := rfl @[to_additive (attr := simp)] -theorem mem_squareIn : a ∈ squareIn G ↔ IsSquare a := Iff.rfl +theorem mem_square {a : G} : a ∈ square G ↔ IsSquare a := Iff.rfl @[to_additive (attr := simp, norm_cast)] -theorem coe_squareIn : squareIn G = {s : G | IsSquare s} := rfl +theorem coe_square : square G = {s : G | IsSquare s} := rfl end Subgroup diff --git a/Mathlib/Algebra/Group/Submonoid/Operations.lean b/Mathlib/Algebra/Group/Submonoid/Operations.lean index a643434cac8a2..8852316886684 100644 --- a/Mathlib/Algebra/Group/Submonoid/Operations.lean +++ b/Mathlib/Algebra/Group/Submonoid/Operations.lean @@ -1000,18 +1000,18 @@ instance smul [SMul M' α] (S : Submonoid M') : SMul S α := @[to_additive] instance smulCommClass_left [SMul M' β] [SMul α β] [SMulCommClass M' α β] (S : Submonoid M') : SMulCommClass S α β := - ⟨fun a _ _ => (smul_comm (a : M') _ _ : _)⟩ + ⟨fun a _ _ => smul_comm (a : M') _ _⟩ @[to_additive] instance smulCommClass_right [SMul α β] [SMul M' β] [SMulCommClass α M' β] (S : Submonoid M') : SMulCommClass α S β := - ⟨fun a s => (smul_comm a (s : M') : _)⟩ + ⟨fun a s => smul_comm a (s : M')⟩ /-- Note that this provides `IsScalarTower S M' M'` which is needed by `SMulMulAssoc`. -/ instance isScalarTower [SMul α β] [SMul M' α] [SMul M' β] [IsScalarTower M' α β] (S : Submonoid M') : IsScalarTower S α β := - ⟨fun a => (smul_assoc (a : M') : _)⟩ + ⟨fun a => smul_assoc (a : M')⟩ section SMul variable [SMul M' α] {S : Submonoid M'} diff --git a/Mathlib/Algebra/Group/UniqueProds/Basic.lean b/Mathlib/Algebra/Group/UniqueProds/Basic.lean index 41d9de70596fd..2e38b8dc940aa 100644 --- a/Mathlib/Algebra/Group/UniqueProds/Basic.lean +++ b/Mathlib/Algebra/Group/UniqueProds/Basic.lean @@ -625,23 +625,6 @@ instance (priority := 100) of_covariant_left [IsLeftCancelMul G] end TwoUniqueProds -@[deprecated (since := "2024-02-04")] -alias UniqueProds.mulHom_image_of_injective := UniqueProds.of_injective_mulHom -@[deprecated (since := "2024-02-04")] -alias UniqueSums.addHom_image_of_injective := UniqueSums.of_injective_addHom -@[deprecated (since := "2024-02-04")] -alias UniqueProds.mulHom_image_iff := MulEquiv.uniqueProds_iff -@[deprecated (since := "2024-02-04")] -alias UniqueSums.addHom_image_iff := AddEquiv.uniqueSums_iff -@[deprecated (since := "2024-02-04")] -alias TwoUniqueProds.mulHom_image_of_injective := TwoUniqueProds.of_injective_mulHom -@[deprecated (since := "2024-02-04")] -alias TwoUniqueSums.addHom_image_of_injective := TwoUniqueSums.of_injective_addHom -@[deprecated (since := "2024-02-04")] -alias TwoUniqueProds.mulHom_image_iff := MulEquiv.twoUniqueProds_iff -@[deprecated (since := "2024-02-04")] -alias TwoUniqueSums.addHom_image_iff := AddEquiv.twoUniqueSums_iff - instance {ι} (G : ι → Type*) [∀ i, AddZeroClass (G i)] [∀ i, TwoUniqueSums (G i)] : TwoUniqueSums (Π₀ i, G i) := TwoUniqueSums.of_injective_addHom diff --git a/Mathlib/Algebra/Group/Units/Basic.lean b/Mathlib/Algebra/Group/Units/Basic.lean index 3630673d90f59..3cb0df004f9d7 100644 --- a/Mathlib/Algebra/Group/Units/Basic.lean +++ b/Mathlib/Algebra/Group/Units/Basic.lean @@ -474,5 +474,3 @@ attribute [deprecated sub_add_cancel_left (since := "2024-03-20")] IsAddUnit.sub -- @[deprecated (since := "2024-03-20")] alias IsUnit.mul_div_cancel := IsUnit.mul_div_cancel_right -- @[deprecated (since := "2024-03-20")] -- alias IsAddUnit.add_sub_cancel := IsAddUnit.add_sub_cancel_right -@[deprecated (since := "2024-03-20")] alias IsUnit.mul_div_cancel' := IsUnit.mul_div_cancel -@[deprecated (since := "2024-03-20")] alias IsAddUnit.add_sub_cancel' := IsAddUnit.add_sub_cancel diff --git a/Mathlib/Algebra/GroupWithZero/Units/Basic.lean b/Mathlib/Algebra/GroupWithZero/Units/Basic.lean index 7702e037e5cb7..5b98cfd3fad68 100644 --- a/Mathlib/Algebra/GroupWithZero/Units/Basic.lean +++ b/Mathlib/Algebra/GroupWithZero/Units/Basic.lean @@ -315,10 +315,6 @@ lemma div_eq_one_iff_eq (hb : b ≠ 0) : a / b = 1 ↔ a = b := hb.isUnit.div_eq lemma div_mul_cancel_right₀ (hb : b ≠ 0) (a : G₀) : b / (a * b) = a⁻¹ := hb.isUnit.div_mul_cancel_right _ -set_option linter.deprecated false in -@[deprecated div_mul_cancel_right₀ (since := "2024-03-20")] -lemma div_mul_left (hb : b ≠ 0) : b / (a * b) = 1 / a := hb.isUnit.div_mul_left - lemma mul_div_mul_right (a b : G₀) (hc : c ≠ 0) : a * c / (b * c) = a / b := hc.isUnit.mul_div_mul_right _ _ @@ -415,10 +411,6 @@ instance (priority := 100) CommGroupWithZero.toDivisionCommMonoid : lemma div_mul_cancel_left₀ (ha : a ≠ 0) (b : G₀) : a / (a * b) = b⁻¹ := ha.isUnit.div_mul_cancel_left _ -@[deprecated div_mul_cancel_left₀ (since := "2024-03-22")] -lemma div_mul_right (b : G₀) (ha : a ≠ 0) : a / (a * b) = 1 / b := by - simp [div_mul_cancel_left₀ ha] - lemma mul_div_cancel_left_of_imp (h : a = 0 → b = 0) : a * b / a = b := by rw [mul_comm, mul_div_cancel_of_imp h] @@ -485,5 +477,3 @@ noncomputable def commGroupWithZeroOfIsUnitOrEqZero [hM : CommMonoidWithZero M] { groupWithZeroOfIsUnitOrEqZero h, hM with } end NoncomputableDefs - -@[deprecated (since := "2024-03-20")] alias mul_div_cancel' := mul_div_cancel₀ diff --git a/Mathlib/Algebra/Homology/Additive.lean b/Mathlib/Algebra/Homology/Additive.lean index 1b14c45bf5b00..86d5e417b4051 100644 --- a/Mathlib/Algebra/Homology/Additive.lean +++ b/Mathlib/Algebra/Homology/Additive.lean @@ -176,7 +176,7 @@ theorem NatTrans.mapHomologicalComplex_naturality {c : ComplexShape ι} {F G : W (α : F ⟶ G) {C D : HomologicalComplex W₁ c} (f : C ⟶ D) : (F.mapHomologicalComplex c).map f ≫ (NatTrans.mapHomologicalComplex α c).app D = (NatTrans.mapHomologicalComplex α c).app C ≫ (G.mapHomologicalComplex c).map f := by - aesop_cat + simp /-- A natural isomorphism between functors induces a natural isomorphism between those functors applied to homological complexes. diff --git a/Mathlib/Algebra/Homology/Augment.lean b/Mathlib/Algebra/Homology/Augment.lean index b7701f75f0d90..2f19bc16f99bb 100644 --- a/Mathlib/Algebra/Homology/Augment.lean +++ b/Mathlib/Algebra/Homology/Augment.lean @@ -37,7 +37,7 @@ The components of this chain map are `C.d 1 0` in degree 0, and zero otherwise. -/ def truncateTo [HasZeroObject V] [HasZeroMorphisms V] (C : ChainComplex V ℕ) : truncate.obj C ⟶ (single₀ V).obj (C.X 0) := - (toSingle₀Equiv (truncate.obj C) (C.X 0)).symm ⟨C.d 1 0, by aesop⟩ + (toSingle₀Equiv (truncate.obj C) (C.X 0)).symm ⟨C.d 1 0, by simp⟩ -- PROJECT when `V` is abelian (but not generally?) -- `[∀ n, Exact (C.d (n+2) (n+1)) (C.d (n+1) n)] [Epi (C.d 1 0)]` iff `QuasiIso (C.truncate_to)` @@ -204,7 +204,7 @@ The components of this chain map are `C.d 0 1` in degree 0, and zero otherwise. -/ def toTruncate [HasZeroObject V] [HasZeroMorphisms V] (C : CochainComplex V ℕ) : (single₀ V).obj (C.X 0) ⟶ truncate.obj C := - (fromSingle₀Equiv (truncate.obj C) (C.X 0)).symm ⟨C.d 0 1, by aesop⟩ + (fromSingle₀Equiv (truncate.obj C) (C.X 0)).symm ⟨C.d 0 1, by simp⟩ variable [HasZeroMorphisms V] diff --git a/Mathlib/Algebra/Homology/HomologicalBicomplex.lean b/Mathlib/Algebra/Homology/HomologicalBicomplex.lean index 634f2df3056f2..e9f7ce59fdfb9 100644 --- a/Mathlib/Algebra/Homology/HomologicalBicomplex.lean +++ b/Mathlib/Algebra/Homology/HomologicalBicomplex.lean @@ -177,7 +177,7 @@ def flipEquivalenceUnitIso : 𝟭 (HomologicalComplex₂ C c₁ c₂) ≅ flipFunctor C c₁ c₂ ⋙ flipFunctor C c₂ c₁ := NatIso.ofComponents (fun K => HomologicalComplex.Hom.isoOfComponents (fun i₁ => HomologicalComplex.Hom.isoOfComponents (fun _ => Iso.refl _) - (by aesop_cat)) (by aesop_cat)) (by aesop_cat) + (by simp)) (by aesop_cat)) (by aesop_cat) /-- Auxiliary definition for `HomologicalComplex₂.flipEquivalence`. -/ @[simps!] @@ -185,7 +185,7 @@ def flipEquivalenceCounitIso : flipFunctor C c₂ c₁ ⋙ flipFunctor C c₁ c₂ ≅ 𝟭 (HomologicalComplex₂ C c₂ c₁) := NatIso.ofComponents (fun K => HomologicalComplex.Hom.isoOfComponents (fun i₂ => HomologicalComplex.Hom.isoOfComponents (fun _ => Iso.refl _) - (by aesop_cat)) (by aesop_cat)) (by aesop_cat) + (by simp)) (by aesop_cat)) (by aesop_cat) /-- Flipping a complex of complexes over the diagonal, as an equivalence of categories. -/ @[simps] diff --git a/Mathlib/Algebra/Homology/HomologySequence.lean b/Mathlib/Algebra/Homology/HomologySequence.lean index 6a325db2f53a9..c3743d891f42f 100644 --- a/Mathlib/Algebra/Homology/HomologySequence.lean +++ b/Mathlib/Algebra/Homology/HomologySequence.lean @@ -165,7 +165,7 @@ 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) + (homologyMap φ j) (by simp) (by simp) (by simp) end HomologySequence @@ -311,7 +311,7 @@ lemma homology_exact₂ : (ShortComplex.mk (HomologicalComplex.homologyMap S.f i have e : S.map (HomologicalComplex.homologyFunctor C c i) ≅ S.map (HomologicalComplex.opcyclesFunctor C c i) := ShortComplex.isoMk (asIso (S.X₁.homologyι i)) - (asIso (S.X₂.homologyι i)) (asIso (S.X₃.homologyι i)) (by aesop_cat) (by aesop_cat) + (asIso (S.X₂.homologyι i)) (asIso (S.X₃.homologyι i)) (by simp) (by simp) exact ShortComplex.exact_of_iso e.symm (opcycles_right_exact S hS.exact i) /-- Exactness of `S.X₂.homology i ⟶ S.X₃.homology i ⟶ S.X₁.homology j`. -/ diff --git a/Mathlib/Algebra/Homology/HomotopyCategory/DegreewiseSplit.lean b/Mathlib/Algebra/Homology/HomotopyCategory/DegreewiseSplit.lean index cabff6665b54e..c454aa0ad7d69 100644 --- a/Mathlib/Algebra/Homology/HomotopyCategory/DegreewiseSplit.lean +++ b/Mathlib/Algebra/Homology/HomotopyCategory/DegreewiseSplit.lean @@ -201,7 +201,7 @@ noncomputable def triangleRotateIsoTriangleOfDegreewiseSplit : (triangle φ).rotate ≅ triangleOfDegreewiseSplit _ (triangleRotateShortComplexSplitting φ) := Triangle.isoMk _ _ (Iso.refl _) (Iso.refl _) (Iso.refl _) - (by aesop_cat) (by aesop_cat) (by ext; dsimp; simp) + (by simp) (by simp) (by ext; dsimp; simp) /-- The triangle `(triangleh φ).rotate` is isomorphic to a triangle attached to a degreewise split short exact sequence of cochain complexes. -/ diff --git a/Mathlib/Algebra/Homology/HomotopyCategory/Pretriangulated.lean b/Mathlib/Algebra/Homology/HomotopyCategory/Pretriangulated.lean index be5caab4950ee..c0a18009abb14 100644 --- a/Mathlib/Algebra/Homology/HomotopyCategory/Pretriangulated.lean +++ b/Mathlib/Algebra/Homology/HomotopyCategory/Pretriangulated.lean @@ -136,8 +136,7 @@ noncomputable def trianglehMapOfHomotopy : comm₃ := by dsimp rw [← Functor.map_comp_assoc, triangleMapOfHomotopy_comm₃, Functor.map_comp, assoc, assoc] - erw [← NatTrans.naturality] - rfl + simp end mapOfHomotopy diff --git a/Mathlib/Algebra/Homology/LocalCohomology.lean b/Mathlib/Algebra/Homology/LocalCohomology.lean index ebc2b550a5d65..7b53d4b469f1a 100644 --- a/Mathlib/Algebra/Homology/LocalCohomology.lean +++ b/Mathlib/Algebra/Homology/LocalCohomology.lean @@ -87,9 +87,7 @@ section variable {R : Type max u v} [CommRing R] {D : Type v} [SmallCategory D] lemma hasColimitDiagram (I : D ⥤ Ideal R) (i : ℕ) : - HasColimit (diagram I i) := by - have : HasColimitsOfShape Dᵒᵖ (AddCommGrpMax.{u, v}) := inferInstance - infer_instance + HasColimit (diagram I i) := inferInstance /- In this definition we do not assume any special property of the diagram `I`, but the relevant case diff --git a/Mathlib/Algebra/Homology/Opposite.lean b/Mathlib/Algebra/Homology/Opposite.lean index e234cce59912e..5bbd8fee09f16 100644 --- a/Mathlib/Algebra/Homology/Opposite.lean +++ b/Mathlib/Algebra/Homology/Opposite.lean @@ -131,9 +131,7 @@ def opUnitIso : 𝟭 (HomologicalComplex V c)ᵒᵖ ≅ opFunctor V c ⋙ opInve intro X Y f refine Quiver.Hom.unop_inj ?_ ext x - simp only [Quiver.Hom.unop_op, Functor.id_map, Iso.op_hom, Functor.comp_map, unop_comp, - comp_f, Hom.isoOfComponents_hom_f] - erw [Category.id_comp, Category.comp_id (f.unop.f x)]) + simp) /-- Auxiliary definition for `opEquivalence`. -/ def opCounitIso : opInverse V c ⋙ opFunctor V c ≅ 𝟭 (HomologicalComplex Vᵒᵖ c.symm) := @@ -182,9 +180,7 @@ def unopUnitIso : 𝟭 (HomologicalComplex Vᵒᵖ c)ᵒᵖ ≅ unopFunctor V c intro X Y f refine Quiver.Hom.unop_inj ?_ ext x - simp only [Quiver.Hom.unop_op, Functor.id_map, Iso.op_hom, Functor.comp_map, unop_comp, - comp_f, Hom.isoOfComponents_hom_f] - erw [Category.id_comp, Category.comp_id (f.unop.f x)]) + simp) /-- Auxiliary definition for `unopEquivalence`. -/ def unopCounitIso : unopInverse V c ⋙ unopFunctor V c ≅ 𝟭 (HomologicalComplex V c.symm) := diff --git a/Mathlib/Algebra/Homology/QuasiIso.lean b/Mathlib/Algebra/Homology/QuasiIso.lean index 47a65f02f7de3..19e311240e6ed 100644 --- a/Mathlib/Algebra/Homology/QuasiIso.lean +++ b/Mathlib/Algebra/Homology/QuasiIso.lean @@ -208,10 +208,8 @@ lemma quasiIso_iff_of_arrow_mk_iso (φ : K ⟶ L) (φ' : K' ⟶ L') (e : Arrow.m [∀ i, K.HasHomology i] [∀ i, L.HasHomology i] [∀ i, K'.HasHomology i] [∀ i, L'.HasHomology i] : QuasiIso φ ↔ QuasiIso φ' := by - rw [← quasiIso_iff_comp_left (show K' ⟶ K from e.inv.left) φ, + simp [← quasiIso_iff_comp_left (show K' ⟶ K from e.inv.left) φ, ← quasiIso_iff_comp_right φ' (show L' ⟶ L from e.inv.right)] - erw [Arrow.w e.inv] - rfl lemma quasiIso_of_arrow_mk_iso (φ : K ⟶ L) (φ' : K' ⟶ L') (e : Arrow.mk φ ≅ Arrow.mk φ') [∀ i, K.HasHomology i] [∀ i, L.HasHomology i] diff --git a/Mathlib/Algebra/Homology/ShortComplex/FunctorEquivalence.lean b/Mathlib/Algebra/Homology/ShortComplex/FunctorEquivalence.lean index 3e5b25181998c..b16976a6c8d8f 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/FunctorEquivalence.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/FunctorEquivalence.lean @@ -50,9 +50,9 @@ def inverse : (J ⥤ ShortComplex C) ⥤ ShortComplex (J ⥤ C) where @[simps!] def unitIso : 𝟭 _ ≅ functor J C ⋙ inverse J C := NatIso.ofComponents (fun _ => isoMk - (NatIso.ofComponents (fun _ => Iso.refl _) (by aesop_cat)) - (NatIso.ofComponents (fun _ => Iso.refl _) (by aesop_cat)) - (NatIso.ofComponents (fun _ => Iso.refl _) (by aesop_cat)) + (NatIso.ofComponents (fun _ => Iso.refl _) (by simp)) + (NatIso.ofComponents (fun _ => Iso.refl _) (by simp)) + (NatIso.ofComponents (fun _ => Iso.refl _) (by simp)) (by aesop_cat) (by aesop_cat)) (by aesop_cat) /-- The counit isomorphism of the equivalence @@ -61,7 +61,7 @@ def unitIso : 𝟭 _ ≅ functor J C ⋙ inverse J C := def counitIso : inverse J C ⋙ functor J C ≅ 𝟭 _ := NatIso.ofComponents (fun _ => NatIso.ofComponents (fun _ => isoMk (Iso.refl _) (Iso.refl _) (Iso.refl _) - (by aesop_cat) (by aesop_cat)) (by aesop_cat)) (by aesop_cat) + (by simp) (by simp)) (by aesop_cat)) (by aesop_cat) end FunctorEquivalence diff --git a/Mathlib/Algebra/Homology/ShortComplex/HomologicalComplex.lean b/Mathlib/Algebra/Homology/ShortComplex/HomologicalComplex.lean index 415cac78ef15f..97bd4dc2b02fa 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/HomologicalComplex.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/HomologicalComplex.lean @@ -50,7 +50,7 @@ when `c.prev j = i` and `c.next j = k`. -/ noncomputable def natIsoSc' (i j k : ι) (hi : c.prev j = i) (hk : c.next j = k) : shortComplexFunctor C c j ≅ shortComplexFunctor' C c i j k := NatIso.ofComponents (fun K => ShortComplex.isoMk (K.XIsoOfEq hi) (Iso.refl _) (K.XIsoOfEq hk) - (by aesop_cat) (by aesop_cat)) (by aesop_cat) + (by simp) (by simp)) (by aesop_cat) variable {C c} diff --git a/Mathlib/Algebra/Homology/ShortComplex/LeftHomology.lean b/Mathlib/Algebra/Homology/ShortComplex/LeftHomology.lean index a8abcedd4972c..8778624055502 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/LeftHomology.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/LeftHomology.lean @@ -1008,7 +1008,7 @@ lemma hasCokernel [S.HasLeftHomology] [HasKernel S.g] : haveI : HasColimit (parallelPair h.f' 0) := ⟨⟨⟨_, h.hπ'⟩⟩⟩ let e : parallelPair (kernel.lift S.g S.f S.zero) 0 ≅ parallelPair h.f' 0 := parallelPair.ext (Iso.refl _) (IsLimit.conePointUniqueUpToIso (kernelIsKernel S.g) h.hi) - (by aesop_cat) (by aesop_cat) + (by aesop_cat) (by simp) exact hasColimitOfIso e end HasLeftHomology diff --git a/Mathlib/Algebra/Homology/ShortComplex/Limits.lean b/Mathlib/Algebra/Homology/ShortComplex/Limits.lean index 1f94e0fc71c5e..a7a46d412c3f3 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/Limits.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/Limits.lean @@ -66,7 +66,7 @@ noncomputable def limitCone : Cone F := Cone.mk (ShortComplex.mk (limMap (whiskerLeft F π₁Toπ₂)) (limMap (whiskerLeft F π₂Toπ₃)) (by aesop_cat)) { app := fun j => Hom.mk (limit.π _ _) (limit.π _ _) (limit.π _ _) - (by aesop_cat) (by aesop_cat) + (by simp) (by simp) naturality := fun _ _ f => by ext all_goals @@ -199,7 +199,7 @@ noncomputable def colimitCocone : Cocone F := Cocone.mk (ShortComplex.mk (colimMap (whiskerLeft F π₁Toπ₂)) (colimMap (whiskerLeft F π₂Toπ₃)) (by aesop_cat)) { app := fun j => Hom.mk (colimit.ι (F ⋙ π₁) _) (colimit.ι (F ⋙ π₂) _) - (colimit.ι (F ⋙ π₃) _) (by aesop_cat) (by aesop_cat) + (colimit.ι (F ⋙ π₃) _) (by simp) (by simp) naturality := fun _ _ f => by ext · dsimp; erw [comp_id, colimit.w (F ⋙ π₁)] diff --git a/Mathlib/Algebra/Homology/ShortComplex/PreservesHomology.lean b/Mathlib/Algebra/Homology/ShortComplex/PreservesHomology.lean index eacf12a85c953..df003423f5f9d 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/PreservesHomology.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/PreservesHomology.lean @@ -834,7 +834,7 @@ lemma preservesLeftHomology_of_zero_g (hg : S.g = 0) f' := by have := h.isIso_i hg let e : parallelPair h.f' 0 ≅ parallelPair S.f 0 := - parallelPair.ext (Iso.refl _) (asIso h.i) (by aesop_cat) (by aesop_cat) + parallelPair.ext (Iso.refl _) (asIso h.i) (by simp) (by simp) exact Limits.preservesColimit_of_iso_diagram F e.symm}⟩ /-- If a short complex `S` is such that `S.f = 0` and that the kernel of `S.g` is preserved @@ -848,7 +848,7 @@ lemma preservesRightHomology_of_zero_f (hf : S.f = 0) g' := by have := h.isIso_p hf let e : parallelPair S.g 0 ≅ parallelPair h.g' 0 := - parallelPair.ext (asIso h.p) (Iso.refl _) (by aesop_cat) (by aesop_cat) + parallelPair.ext (asIso h.p) (Iso.refl _) (by simp) (by simp) exact Limits.preservesLimit_of_iso_diagram F e }⟩ end Functor diff --git a/Mathlib/Algebra/Homology/ShortComplex/RightHomology.lean b/Mathlib/Algebra/Homology/ShortComplex/RightHomology.lean index 9832970621a74..b6c4767c9e7b7 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/RightHomology.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/RightHomology.lean @@ -150,7 +150,7 @@ def ofIsLimitKernelFork (hf : S.f = 0) (c : KernelFork S.g) (hc : IsLimit c) : wp := by rw [comp_id, hf] hp := CokernelCofork.IsColimit.ofId _ hf wι := KernelFork.condition _ - hι := IsLimit.ofIsoLimit hc (Fork.ext (Iso.refl _) (by aesop_cat)) + hι := IsLimit.ofIsoLimit hc (Fork.ext (Iso.refl _) (by simp)) @[simp] lemma ofIsLimitKernelFork_g' (hf : S.f = 0) (c : KernelFork S.g) (hc : IsLimit c) : (ofIsLimitKernelFork S hf c hc).g' = S.g := by @@ -173,7 +173,7 @@ def ofIsColimitCokernelCofork (hg : S.g = 0) (c : CokernelCofork S.f) (hc : IsCo p := c.π ι := 𝟙 _ wp := CokernelCofork.condition _ - hp := IsColimit.ofIsoColimit hc (Cofork.ext (Iso.refl _) (by aesop_cat)) + hp := IsColimit.ofIsoColimit hc (Cofork.ext (Iso.refl _) (by simp)) wι := Cofork.IsColimit.hom_ext hc (by simp [hg]) hι := KernelFork.IsLimit.ofId _ (Cofork.IsColimit.hom_ext hc (by simp [hg])) @@ -1287,7 +1287,7 @@ lemma hasKernel [S.HasRightHomology] [HasCokernel S.f] : haveI : HasLimit (parallelPair h.g' 0) := ⟨⟨⟨_, h.hι'⟩⟩⟩ let e : parallelPair (cokernel.desc S.f S.g S.zero) 0 ≅ parallelPair h.g' 0 := parallelPair.ext (IsColimit.coconePointUniqueUpToIso (colimit.isColimit _) h.hp) - (Iso.refl _) (coequalizer.hom_ext (by simp)) (by aesop_cat) + (Iso.refl _) (coequalizer.hom_ext (by simp)) (by simp) exact hasLimitOfIso e.symm end HasRightHomology diff --git a/Mathlib/Algebra/Homology/ShortComplex/SnakeLemma.lean b/Mathlib/Algebra/Homology/ShortComplex/SnakeLemma.lean index 8fe3dbac06932..c3abf86cbb643 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/SnakeLemma.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/SnakeLemma.lean @@ -345,7 +345,7 @@ lemma op_δ : S.op.δ = S.δ.op := Quiver.Hom.unop_inj (by /-- The duality isomorphism `S.L₂'.op ≅ S.op.L₁'`. -/ noncomputable def L₂'OpIso : S.L₂'.op ≅ S.op.L₁' := - ShortComplex.isoMk (Iso.refl _) (Iso.refl _) (Iso.refl _) (by aesop_cat) + ShortComplex.isoMk (Iso.refl _) (Iso.refl _) (Iso.refl _) (by simp) (by dsimp; simp only [id_comp, comp_id, S.op_δ]) /-- Exactness of `L₀.X₃ ⟶ L₃.X₁ ⟶ L₃.X₂`. -/ @@ -475,7 +475,7 @@ noncomputable def functorP : SnakeInput C ⥤ C where obj S := S.P map f := pullback.map _ _ _ _ f.f₁.τ₂ f.f₀.τ₃ f.f₁.τ₃ f.f₁.comm₂₃.symm (congr_arg ShortComplex.Hom.τ₃ f.comm₀₁.symm) - map_id _ := by dsimp [P]; aesop_cat + map_id _ := by dsimp [P]; simp map_comp _ _ := by dsimp [P]; aesop_cat @[reassoc] diff --git a/Mathlib/Algebra/Homology/Single.lean b/Mathlib/Algebra/Homology/Single.lean index a9cbef23cd3c9..b6dc763c87a95 100644 --- a/Mathlib/Algebra/Homology/Single.lean +++ b/Mathlib/Algebra/Homology/Single.lean @@ -220,7 +220,7 @@ noncomputable def toSingle₀Equiv (C : ChainComplex V ℕ) (X : V) : obtain rfl : i = 1 := by simpa using hi.symm exact f.2) left_inv φ := by aesop_cat - right_inv f := by aesop_cat + right_inv f := by simp @[simp] lemma toSingle₀Equiv_symm_apply_f_zero {C : ChainComplex V ℕ} {X : V} diff --git a/Mathlib/Algebra/Homology/TotalComplexShift.lean b/Mathlib/Algebra/Homology/TotalComplexShift.lean index 5c418bdf20985..c1ab4298eb34f 100644 --- a/Mathlib/Algebra/Homology/TotalComplexShift.lean +++ b/Mathlib/Algebra/Homology/TotalComplexShift.lean @@ -102,18 +102,8 @@ instance : ((shiftFunctor₂ C y).obj K).HasTotal (up ℤ) := fun n => invFun := fun ⟨⟨a, b⟩, h⟩ => ⟨(a, b - y), by simp only [Set.mem_preimage, instTotalComplexShape_π, Set.mem_singleton_iff] at h ⊢ omega⟩ - left_inv := by - rintro ⟨⟨a, b⟩, h⟩ - ext - · rfl - · dsimp - omega - right_inv := by - intro ⟨⟨a, b⟩, h⟩ - ext - · rfl - · dsimp - omega } + left_inv _ := by simp + right_inv _ := by simp } (fun _ => Iso.refl _) instance : ((shiftFunctor₂ C y ⋙ shiftFunctor₁ C x).obj K).HasTotal (up ℤ) := by diff --git a/Mathlib/Algebra/Lie/BaseChange.lean b/Mathlib/Algebra/Lie/BaseChange.lean index d6be0876fe3d9..86e3b96fc02af 100644 --- a/Mathlib/Algebra/Lie/BaseChange.lean +++ b/Mathlib/Algebra/Lie/BaseChange.lean @@ -135,7 +135,7 @@ variable [CommRing A] [LieAlgebra A L] instance lieAlgebra [CommRing R] [Algebra R A] : LieAlgebra R (RestrictScalars R A L) where lie_smul t x y := (lie_smul (algebraMap R A t) (RestrictScalars.addEquiv R A L x) - (RestrictScalars.addEquiv R A L y) : _) + (RestrictScalars.addEquiv R A L y) :) end RestrictScalars diff --git a/Mathlib/Algebra/Lie/Basic.lean b/Mathlib/Algebra/Lie/Basic.lean index 8c75a9021fe88..7a022b824e648 100644 --- a/Mathlib/Algebra/Lie/Basic.lean +++ b/Mathlib/Algebra/Lie/Basic.lean @@ -200,13 +200,11 @@ theorem lie_nsmul (n : ℕ) : ⁅x, n • m⁆ = n • ⁅x, m⁆ := { toFun := fun m : M => ⁅x, m⁆, map_zero' := lie_zero x, map_add' := fun _ _ => lie_add _ _ _} _ _ -@[simp] theorem zsmul_lie (a : ℤ) : ⁅a • x, m⁆ = a • ⁅x, m⁆ := AddMonoidHom.map_zsmul { toFun := fun x : L => ⁅x, m⁆, map_zero' := zero_lie m, map_add' := fun _ _ => add_lie _ _ _ } _ _ -@[simp] theorem lie_zsmul (a : ℤ) : ⁅x, a • m⁆ = a • ⁅x, m⁆ := AddMonoidHom.map_zsmul { toFun := fun m : M => ⁅x, m⁆, map_zero' := lie_zero x, map_add' := fun _ _ => lie_add _ _ _ } @@ -221,6 +219,10 @@ theorem lie_jacobi : ⁅x, ⁅y, z⁆⁆ + ⁅y, ⁅z, x⁆⁆ + ⁅z, ⁅x, y instance LieRing.instLieAlgebra : LieAlgebra ℤ L where lie_smul n x y := lie_zsmul x y n +instance : LieModule ℤ L M where + smul_lie n x m := zsmul_lie x m n + lie_smul n x m := lie_zsmul x m n + instance LinearMap.instLieRingModule : LieRingModule L (M →ₗ[R] N) where bracket x f := { toFun := fun m => ⁅x, f m⁆ - f ⁅x, m⁆ diff --git a/Mathlib/Algebra/Lie/CartanSubalgebra.lean b/Mathlib/Algebra/Lie/CartanSubalgebra.lean index cfe41bd7c8bfd..1887a29efd5ca 100644 --- a/Mathlib/Algebra/Lie/CartanSubalgebra.lean +++ b/Mathlib/Algebra/Lie/CartanSubalgebra.lean @@ -44,10 +44,10 @@ namespace LieSubalgebra A _splitting_ Cartan subalgebra can be defined by mixing in `LieModule.IsTriangularizable R H L`. -/ class IsCartanSubalgebra : Prop where - nilpotent : LieAlgebra.IsNilpotent R H + nilpotent : LieRing.IsNilpotent H self_normalizing : H.normalizer = H -instance [H.IsCartanSubalgebra] : LieAlgebra.IsNilpotent R H := +instance [H.IsCartanSubalgebra] : LieRing.IsNilpotent H := IsCartanSubalgebra.nilpotent @[simp] @@ -66,7 +66,7 @@ theorem ucs_eq_self_of_isCartanSubalgebra (H : LieSubalgebra R L) [H.IsCartanSub theorem isCartanSubalgebra_iff_isUcsLimit : H.IsCartanSubalgebra ↔ H.toLieSubmodule.IsUcsLimit := by constructor · intro h - have h₁ : LieAlgebra.IsNilpotent R H := by infer_instance + have h₁ : LieRing.IsNilpotent H := by infer_instance obtain ⟨k, hk⟩ := H.toLieSubmodule.isNilpotent_iff_exists_self_le_ucs.mp h₁ replace hk : H.toLieSubmodule = LieSubmodule.ucs k ⊥ := le_antisymm hk @@ -77,7 +77,7 @@ theorem isCartanSubalgebra_iff_isUcsLimit : H.IsCartanSubalgebra ↔ H.toLieSubm · rintro ⟨k, hk⟩ exact { nilpotent := by - dsimp only [LieAlgebra.IsNilpotent] + dsimp only [LieRing.IsNilpotent] erw [H.toLieSubmodule.isNilpotent_iff_exists_lcs_eq_bot] use k rw [_root_.eq_bot_iff, LieSubmodule.lcs_le_iff, hk k (le_refl k)] @@ -113,7 +113,7 @@ theorem LieIdeal.normalizer_eq_top {R : Type u} {L : Type v} [CommRing R] [LieRi open LieIdeal /-- A nilpotent Lie algebra is its own Cartan subalgebra. -/ -instance LieAlgebra.top_isCartanSubalgebra_of_nilpotent [LieAlgebra.IsNilpotent R L] : +instance LieAlgebra.top_isCartanSubalgebra_of_nilpotent [LieRing.IsNilpotent L] : LieSubalgebra.IsCartanSubalgebra (⊤ : LieSubalgebra R L) where nilpotent := inferInstance self_normalizing := by rw [← top_toLieSubalgebra, normalizer_eq_top, top_toLieSubalgebra] diff --git a/Mathlib/Algebra/Lie/Engel.lean b/Mathlib/Algebra/Lie/Engel.lean index 4380e2dd3aef3..1ce0e3176e8a1 100644 --- a/Mathlib/Algebra/Lie/Engel.lean +++ b/Mathlib/Algebra/Lie/Engel.lean @@ -124,12 +124,13 @@ theorem lcs_le_lcs_of_is_nilpotent_span_sup_eq_top {n i j : ℕ} exact antitone_lowerCentralSeries R L M le_self_add theorem isNilpotentOfIsNilpotentSpanSupEqTop (hnp : IsNilpotent <| toEnd R L M x) - (hIM : IsNilpotent R I M) : IsNilpotent R L M := by + (hIM : IsNilpotent I M) : IsNilpotent L M := by obtain ⟨n, hn⟩ := hnp - obtain ⟨k, hk⟩ := hIM + obtain ⟨k, hk⟩ := IsNilpotent.nilpotent R I M have hk' : I.lcs M k = ⊥ := by simp only [← toSubmodule_inj, I.coe_lcs_eq, hk, bot_toSubmodule] suffices ∀ l, lowerCentralSeries R L M (l * n) ≤ I.lcs M l by + rw [isNilpotent_iff R] use k * n simpa [hk'] using this k intro l @@ -154,15 +155,14 @@ Engel's theorem `LieAlgebra.isEngelian_of_isNoetherian` states that any Noetheri Engelian. -/ def LieAlgebra.IsEngelian : Prop := ∀ (M : Type u₄) [AddCommGroup M] [Module R M] [LieRingModule L M] [LieModule R L M], - (∀ x : L, _root_.IsNilpotent (toEnd R L M x)) → LieModule.IsNilpotent R L M + (∀ x : L, _root_.IsNilpotent (toEnd R L M x)) → LieModule.IsNilpotent L M variable {R L} theorem LieAlgebra.isEngelian_of_subsingleton [Subsingleton L] : LieAlgebra.IsEngelian R L := by intro M _i1 _i2 _i3 _i4 _h use 1 - suffices (⊤ : LieIdeal R L) = ⊥ by simp [this] - subsingleton [(LieSubmodule.subsingleton_iff R L L).mpr inferInstance] + simp theorem Function.Surjective.isEngelian {f : L →ₗ⁅R⁆ L₂} (hf : Function.Surjective f) (h : LieAlgebra.IsEngelian.{u₁, u₂, u₄} R L) : LieAlgebra.IsEngelian.{u₁, u₃, u₄} R L₂ := by @@ -171,7 +171,7 @@ theorem Function.Surjective.isEngelian {f : L →ₗ⁅R⁆ L₂} (hf : Function letI : LieModule R L M := compLieHom M f have hnp : ∀ x, IsNilpotent (toEnd R L M x) := fun x => h' (f x) have surj_id : Function.Surjective (LinearMap.id : M →ₗ[R] M) := Function.surjective_id - haveI : LieModule.IsNilpotent R L M := h M hnp + haveI : LieModule.IsNilpotent L M := h M hnp apply hf.lieModuleIsNilpotent surj_id -- Porting note (https://github.com/leanprover-community/mathlib4/issues/10745): was `simp` intros; simp only [LinearMap.id_coe, id_eq]; rfl @@ -218,16 +218,16 @@ Note that this implies all traditional forms of Engel's theorem via `LieAlgebra.isNilpotent_iff_forall`. -/ theorem LieAlgebra.isEngelian_of_isNoetherian [IsNoetherian R L] : LieAlgebra.IsEngelian R L := by intro M _i1 _i2 _i3 _i4 h - rw [← isNilpotent_range_toEnd_iff] + rw [← isNilpotent_range_toEnd_iff R] let L' := (toEnd R L M).range replace h : ∀ y : L', _root_.IsNilpotent (y : Module.End R M) := by rintro ⟨-, ⟨y, rfl⟩⟩ simp [h] - change LieModule.IsNilpotent R L' M + change LieModule.IsNilpotent L' M let s := {K : LieSubalgebra R L' | LieAlgebra.IsEngelian R K} have hs : s.Nonempty := ⟨⊥, LieAlgebra.isEngelian_of_subsingleton⟩ suffices ⊤ ∈ s by - rw [← isNilpotent_of_top_iff] + rw [← isNilpotent_of_top_iff (R := R)] apply this M simp [LieSubalgebra.toEnd_eq, h] have : ∀ K ∈ s, K ≠ ⊤ → ∃ K' ∈ s, K < K' := by @@ -242,7 +242,7 @@ theorem LieAlgebra.isEngelian_of_isNoetherian [IsNoetherian R L] : LieAlgebra.Is LieSubmodule.top_toSubmodule, ← LieSubalgebra.top_toSubmodule, K.toSubmodule_inj] exact Submodule.Quotient.nontrivial_of_lt_top _ hK₂.lt_top - have : LieModule.IsNilpotent R K (L' ⧸ K.toLieSubmodule) := by + have : LieModule.IsNilpotent K (L' ⧸ K.toLieSubmodule) := by -- Porting note: was refine' hK₁ _ fun x => _ apply hK₁ intro x @@ -274,18 +274,18 @@ theorem LieAlgebra.isEngelian_of_isNoetherian [IsNoetherian R L] : LieAlgebra.Is See also `LieModule.isNilpotent_iff_forall'` which assumes that `M` is Noetherian instead of `L`. -/ theorem LieModule.isNilpotent_iff_forall [IsNoetherian R L] : - LieModule.IsNilpotent R L M ↔ ∀ x, _root_.IsNilpotent <| toEnd R L M x := + LieModule.IsNilpotent L M ↔ ∀ x, _root_.IsNilpotent <| toEnd R L M x := ⟨fun _ ↦ isNilpotent_toEnd_of_isNilpotent R L M, fun h => LieAlgebra.isEngelian_of_isNoetherian M h⟩ /-- Engel's theorem. -/ theorem LieModule.isNilpotent_iff_forall' [IsNoetherian R M] : - LieModule.IsNilpotent R L M ↔ ∀ x, _root_.IsNilpotent <| toEnd R L M x := by - rw [← isNilpotent_range_toEnd_iff, LieModule.isNilpotent_iff_forall]; simp + LieModule.IsNilpotent L M ↔ ∀ x, _root_.IsNilpotent <| toEnd R L M x := by + rw [← isNilpotent_range_toEnd_iff (R := R), LieModule.isNilpotent_iff_forall (R := R)]; simp /-- Engel's theorem. -/ theorem LieAlgebra.isNilpotent_iff_forall [IsNoetherian R L] : - LieAlgebra.IsNilpotent R L ↔ ∀ x, _root_.IsNilpotent <| LieAlgebra.ad R L x := + LieRing.IsNilpotent L ↔ ∀ x, _root_.IsNilpotent <| LieAlgebra.ad R L x := LieModule.isNilpotent_iff_forall end LieAlgebra diff --git a/Mathlib/Algebra/Lie/EngelSubalgebra.lean b/Mathlib/Algebra/Lie/EngelSubalgebra.lean index e011fcd6e98b4..65e5d19fc51f3 100644 --- a/Mathlib/Algebra/Lie/EngelSubalgebra.lean +++ b/Mathlib/Algebra/Lie/EngelSubalgebra.lean @@ -143,8 +143,8 @@ lemma normalizer_eq_self_of_engel_le [IsArtinian R L] if it is contained in the Engel subalgebra of all its elements. -/ lemma isNilpotent_of_forall_le_engel [IsNoetherian R L] (H : LieSubalgebra R L) (h : ∀ x ∈ H, H ≤ engel R x) : - LieAlgebra.IsNilpotent R H := by - rw [LieAlgebra.isNilpotent_iff_forall] + LieRing.IsNilpotent H := by + rw [LieAlgebra.isNilpotent_iff_forall (R := R)] intro x let K : ℕ →o Submodule R H := ⟨fun n ↦ LinearMap.ker ((ad R H x) ^ n), fun m n hmn ↦ ?mono⟩ diff --git a/Mathlib/Algebra/Lie/IdealOperations.lean b/Mathlib/Algebra/Lie/IdealOperations.lean index aae8ed34539a5..73180718964b1 100644 --- a/Mathlib/Algebra/Lie/IdealOperations.lean +++ b/Mathlib/Algebra/Lie/IdealOperations.lean @@ -69,19 +69,18 @@ variable [LieAlgebra R L] [LieModule R L M₂] (I J : LieIdeal R L) /-- Given a Lie module `M` over a Lie algebra `L`, the set of Lie ideals of `L` acts on the set of submodules of `M`. -/ instance hasBracket : Bracket (LieIdeal R L) (LieSubmodule R L M) := - ⟨fun I N => lieSpan R L { m | ∃ (x : I) (n : N), ⁅(x : L), (n : M)⁆ = m }⟩ + ⟨fun I N => lieSpan R L { ⁅(x : L), (n : M)⁆ | (x : I) (n : N) }⟩ theorem lieIdeal_oper_eq_span : - ⁅I, N⁆ = lieSpan R L { m | ∃ (x : I) (n : N), ⁅(x : L), (n : M)⁆ = m } := + ⁅I, N⁆ = lieSpan R L { ⁅(x : L), (n : M)⁆ | (x : I) (n : N) } := rfl /-- See also `LieSubmodule.lieIdeal_oper_eq_linear_span'` and `LieSubmodule.lieIdeal_oper_eq_tensor_map_range`. -/ theorem lieIdeal_oper_eq_linear_span [LieModule R L M] : - (↑⁅I, N⁆ : Submodule R M) = - Submodule.span R { m | ∃ (x : I) (n : N), ⁅(x : L), (n : M)⁆ = m } := by + (↑⁅I, N⁆ : Submodule R M) = Submodule.span R { ⁅(x : L), (n : M)⁆ | (x : I) (n : N) } := by apply le_antisymm - · let s := { m : M | ∃ (x : ↥I) (n : ↥N), ⁅(x : L), (n : M)⁆ = m } + · let s := { ⁅(x : L), (n : M)⁆ | (x : I) (n : N) } have aux : ∀ (y : L), ∀ m' ∈ Submodule.span R s, ⁅y, m'⁆ ∈ Submodule.span R s := by intro y m' hm' refine Submodule.span_induction (R := R) (M := M) (s := s) @@ -99,7 +98,7 @@ theorem lieIdeal_oper_eq_linear_span [LieModule R L M] : · rw [lieIdeal_oper_eq_span]; apply submodule_span_le_lieSpan theorem lieIdeal_oper_eq_linear_span' [LieModule R L M] : - (↑⁅I, N⁆ : Submodule R M) = Submodule.span R { m | ∃ x ∈ I, ∃ n ∈ N, ⁅x, n⁆ = m } := by + (↑⁅I, N⁆ : Submodule R M) = Submodule.span R { ⁅x, n⁆ | (x ∈ I) (n ∈ N) } := by rw [lieIdeal_oper_eq_linear_span] congr ext m diff --git a/Mathlib/Algebra/Lie/LieTheorem.lean b/Mathlib/Algebra/Lie/LieTheorem.lean index a9e21ea679bd7..b9df89d846f10 100644 --- a/Mathlib/Algebra/Lie/LieTheorem.lean +++ b/Mathlib/Algebra/Lie/LieTheorem.lean @@ -204,7 +204,7 @@ open LieAlgebra -- except that it additionally assumes a finiteness hypothesis. private lemma exists_forall_lie_eq_smul_of_isSolvable_of_finite (L : Type*) [LieRing L] [LieAlgebra k L] [LieRingModule L V] [LieModule k L V] - [IsSolvable k L] [LieModule.IsTriangularizable k L V] [Module.Finite k L] : + [IsSolvable L] [LieModule.IsTriangularizable k L V] [Module.Finite k L] : ∃ χ : Module.Dual k L, Nontrivial (weightSpace V χ) := by obtain H|⟨A, hA, hAL⟩ := eq_top_or_exists_le_coatom (derivedSeries k L 1).toSubmodule · obtain _|_ := subsingleton_or_nontrivial L @@ -231,7 +231,7 @@ have a common eigenvector for the action of all elements of the Lie algebra. See `LieModule.exists_nontrivial_weightSpace_of_isNilpotent` for the variant that assumes that `L` is nilpotent and drops the condition that `k` is of characteristic zero. -/ theorem exists_nontrivial_weightSpace_of_isSolvable - [IsSolvable k L] [LieModule.IsTriangularizable k L V] : + [IsSolvable L] [LieModule.IsTriangularizable k L V] : ∃ χ : Module.Dual k L, Nontrivial (weightSpace V χ) := by let imL := (toEnd k L V).range let toEndo : L →ₗ[k] imL := LinearMap.codRestrict imL.toSubmodule (toEnd k L V) diff --git a/Mathlib/Algebra/Lie/Nilpotent.lean b/Mathlib/Algebra/Lie/Nilpotent.lean index 7abb7b337eb4f..4e19fe0b77a4d 100644 --- a/Mathlib/Algebra/Lie/Nilpotent.lean +++ b/Mathlib/Algebra/Lie/Nilpotent.lean @@ -88,6 +88,39 @@ theorem lowerCentralSeries_succ : lowerCentralSeries R L M (k + 1) = ⁅(⊤ : LieIdeal R L), lowerCentralSeries R L M k⁆ := (⊤ : LieSubmodule R L M).lcs_succ k +private theorem coe_lowerCentralSeries_eq_int_aux (R₁ R₂ L M : Type*) + [CommRing R₁] [CommRing R₂] [AddCommGroup M] + [LieRing L] [LieAlgebra R₁ L] [LieAlgebra R₂ L] [Module R₁ M] [Module R₂ M] [LieRingModule L M] + [LieModule R₁ L M] (k : ℕ) : + let I := lowerCentralSeries R₂ L M k; let S : Set M := {⁅a, b⁆ | (a : L) (b ∈ I)} + (Submodule.span R₁ S : Set M) ≤ (Submodule.span R₂ S : Set M) := by + intro I S x hx + simp only [SetLike.mem_coe] at hx ⊢ + induction hx using Submodule.closure_induction with + | zero => exact Submodule.zero_mem _ + | add y z hy₁ hz₁ hy₂ hz₂ => exact Submodule.add_mem _ hy₂ hz₂ + | smul_mem c y hy => + obtain ⟨a, b, hb, rfl⟩ := hy + rw [← smul_lie] + exact Submodule.subset_span ⟨c • a, b, hb, rfl⟩ + +theorem coe_lowerCentralSeries_eq_int [LieModule R L M] (k : ℕ) : + (lowerCentralSeries R L M k : Set M) = (lowerCentralSeries ℤ L M k : Set M) := by + show ((lowerCentralSeries R L M k).toSubmodule : Set M) = + ((lowerCentralSeries ℤ L M k).toSubmodule : Set M) + induction k with + | zero => rfl + | succ k ih => + rw [lowerCentralSeries_succ, lowerCentralSeries_succ] + rw [LieSubmodule.lieIdeal_oper_eq_linear_span', LieSubmodule.lieIdeal_oper_eq_linear_span'] + rw [Set.ext_iff] at ih + simp only [SetLike.mem_coe, LieSubmodule.mem_toSubmodule] at ih + simp only [LieSubmodule.mem_top, ih, true_and] + apply le_antisymm + · exact coe_lowerCentralSeries_eq_int_aux _ _ L M k + · simp only [← ih] + exact coe_lowerCentralSeries_eq_int_aux _ _ L M k + end LieModule namespace LieSubmodule @@ -148,7 +181,7 @@ theorem eventually_iInf_lowerCentralSeries_eq [IsArtinian R M] : theorem trivial_iff_lower_central_eq_bot : IsTrivial L M ↔ lowerCentralSeries R L M 1 = ⊥ := by constructor <;> intro h - · erw [eq_bot_iff, LieSubmodule.lieSpan_le]; rintro m ⟨x, n, hn⟩; rw [← hn, h.trivial]; simp + · simp · rw [LieSubmodule.eq_bot_iff] at h; apply IsTrivial.mk; intro x m; apply h apply LieSubmodule.subset_lieSpan -- Porting note: was `use x, m; rfl` @@ -217,56 +250,73 @@ theorem derivedSeries_le_lowerCentralSeries (k : ℕ) : /-- A Lie module is nilpotent if its lower central series reaches 0 (in a finite number of steps). -/ +@[mk_iff isNilpotent_iff_int] class IsNilpotent : Prop where - nilpotent : ∃ k, lowerCentralSeries R L M k = ⊥ + mk_int :: + nilpotent_int : ∃ k, lowerCentralSeries ℤ L M k = ⊥ + +section + +variable [LieModule R L M] + +/-- See also `LieModule.isNilpotent_iff_exists_ucs_eq_top`. -/ +lemma isNilpotent_iff : + IsNilpotent L M ↔ ∃ k, lowerCentralSeries R L M k = ⊥ := by + simp [isNilpotent_iff_int, SetLike.ext'_iff, coe_lowerCentralSeries_eq_int R L M] -theorem exists_lowerCentralSeries_eq_bot_of_isNilpotent [IsNilpotent R L M] : +lemma IsNilpotent.nilpotent [IsNilpotent L M] : ∃ k, lowerCentralSeries R L M k = ⊥ := + (isNilpotent_iff R L M).mp ‹_› + +variable {R L} in +lemma IsNilpotent.mk {k : ℕ} (h : lowerCentralSeries R L M k = ⊥) : IsNilpotent L M := + (isNilpotent_iff R L M).mpr ⟨k, h⟩ + +@[deprecated IsNilpotent.nilpotent (since := "2025-01-07")] +theorem exists_lowerCentralSeries_eq_bot_of_isNilpotent [IsNilpotent L M] : ∃ k, lowerCentralSeries R L M k = ⊥ := - IsNilpotent.nilpotent + IsNilpotent.nilpotent R L M -@[simp] lemma iInf_lowerCentralSeries_eq_bot_of_isNilpotent [IsNilpotent R L M] : +@[simp] lemma iInf_lowerCentralSeries_eq_bot_of_isNilpotent [IsNilpotent L M] : ⨅ k, lowerCentralSeries R L M k = ⊥ := by - obtain ⟨k, hk⟩ := exists_lowerCentralSeries_eq_bot_of_isNilpotent R L M + obtain ⟨k, hk⟩ := IsNilpotent.nilpotent R L M rw [eq_bot_iff, ← hk] exact iInf_le _ _ -/-- See also `LieModule.isNilpotent_iff_exists_ucs_eq_top`. -/ -theorem isNilpotent_iff : IsNilpotent R L M ↔ ∃ k, lowerCentralSeries R L M k = ⊥ := - ⟨fun h => h.nilpotent, fun h => ⟨h⟩⟩ +end section variable {R L M} variable [LieModule R L M] theorem _root_.LieSubmodule.isNilpotent_iff_exists_lcs_eq_bot (N : LieSubmodule R L M) : - LieModule.IsNilpotent R L N ↔ ∃ k, N.lcs k = ⊥ := by - rw [isNilpotent_iff] + LieModule.IsNilpotent L N ↔ ∃ k, N.lcs k = ⊥ := by + rw [isNilpotent_iff R L N] refine exists_congr fun k => ?_ rw [N.lowerCentralSeries_eq_lcs_comap k, LieSubmodule.comap_incl_eq_bot, inf_eq_right.mpr (N.lcs_le_self k)] variable (R L M) -instance (priority := 100) trivialIsNilpotent [IsTrivial L M] : IsNilpotent R L M := +instance (priority := 100) trivialIsNilpotent [IsTrivial L M] : IsNilpotent L M := ⟨by use 1; change ⁅⊤, ⊤⁆ = ⊥; simp⟩ -theorem exists_forall_pow_toEnd_eq_zero [hM : IsNilpotent R L M] : +theorem exists_forall_pow_toEnd_eq_zero [IsNilpotent L M] : ∃ k : ℕ, ∀ x : L, toEnd R L M x ^ k = 0 := by - obtain ⟨k, hM⟩ := hM + obtain ⟨k, hM⟩ := IsNilpotent.nilpotent R L M use k intro x; ext m rw [LinearMap.pow_apply, LinearMap.zero_apply, ← @LieSubmodule.mem_bot R L M, ← hM] exact iterate_toEnd_mem_lowerCentralSeries R L M x m k -theorem isNilpotent_toEnd_of_isNilpotent [IsNilpotent R L M] (x : L) : +theorem isNilpotent_toEnd_of_isNilpotent [IsNilpotent L M] (x : L) : _root_.IsNilpotent (toEnd R L M x) := by change ∃ k, toEnd R L M x ^ k = 0 have := exists_forall_pow_toEnd_eq_zero R L M tauto -theorem isNilpotent_toEnd_of_isNilpotent₂ [IsNilpotent R L M] (x y : L) : +theorem isNilpotent_toEnd_of_isNilpotent₂ [IsNilpotent L M] (x y : L) : _root_.IsNilpotent (toEnd R L M x ∘ₗ toEnd R L M y) := by - obtain ⟨k, hM⟩ := exists_lowerCentralSeries_eq_bot_of_isNilpotent R L M + obtain ⟨k, hM⟩ := IsNilpotent.nilpotent R L M replace hM : lowerCentralSeries R L M (2 * k) = ⊥ := by rw [eq_bot_iff, ← hM]; exact antitone_lowerCentralSeries R L M (by omega) use k @@ -274,7 +324,7 @@ theorem isNilpotent_toEnd_of_isNilpotent₂ [IsNilpotent R L M] (x y : L) : rw [LinearMap.pow_apply, LinearMap.zero_apply, ← LieSubmodule.mem_bot (R := R) (L := L), ← hM] exact iterate_toEnd_mem_lowerCentralSeries₂ R L M x y m k -@[simp] lemma maxGenEigenSpace_toEnd_eq_top [IsNilpotent R L M] (x : L) : +@[simp] lemma maxGenEigenSpace_toEnd_eq_top [IsNilpotent L M] (x : L) : ((toEnd R L M x).maxGenEigenspace 0) = ⊤ := by ext m simp only [Module.End.mem_maxGenEigenspace, zero_smul, sub_zero, Submodule.mem_top, @@ -289,7 +339,8 @@ This is essentially the Lie module equivalent of the fact that a central extension of nilpotent Lie algebras is nilpotent. See `LieAlgebra.nilpotent_of_nilpotent_quotient` below for the corresponding result for Lie algebras. -/ theorem nilpotentOfNilpotentQuotient {N : LieSubmodule R L M} (h₁ : N ≤ maxTrivSubmodule R L M) - (h₂ : IsNilpotent R L (M ⧸ N)) : IsNilpotent R L M := by + (h₂ : IsNilpotent L (M ⧸ N)) : IsNilpotent L M := by + rw [isNilpotent_iff R L] at h₂ ⊢ obtain ⟨k, hk⟩ := h₂ use k + 1 simp only [lowerCentralSeries_succ] @@ -300,13 +351,13 @@ theorem nilpotentOfNilpotentQuotient {N : LieSubmodule R L M} (h₁ : N ≤ maxT exact map_lowerCentralSeries_le k (LieSubmodule.Quotient.mk' N) theorem isNilpotent_quotient_iff : - IsNilpotent R L (M ⧸ N) ↔ ∃ k, lowerCentralSeries R L M k ≤ N := by - rw [LieModule.isNilpotent_iff] + IsNilpotent L (M ⧸ N) ↔ ∃ k, lowerCentralSeries R L M k ≤ N := by + rw [isNilpotent_iff R L] refine exists_congr fun k ↦ ?_ rw [← LieSubmodule.Quotient.map_mk'_eq_bot_le, map_lowerCentralSeries_eq k (LieSubmodule.Quotient.surjective_mk' N)] -theorem iInf_lcs_le_of_isNilpotent_quot (h : IsNilpotent R L (M ⧸ N)) : +theorem iInf_lcs_le_of_isNilpotent_quot (h : IsNilpotent L (M ⧸ N)) : ⨅ k, lowerCentralSeries R L M k ≤ N := by obtain ⟨k, hk⟩ := (isNilpotent_quotient_iff R L M N).mp h exact iInf_le_of_le k hk @@ -318,81 +369,90 @@ the natural number `k` (the number of inclusions). For a non-nilpotent module, we use the junk value 0. -/ noncomputable def nilpotencyLength : ℕ := - sInf {k | lowerCentralSeries R L M k = ⊥} + sInf {k | lowerCentralSeries ℤ L M k = ⊥} @[simp] -theorem nilpotencyLength_eq_zero_iff [IsNilpotent R L M] : - nilpotencyLength R L M = 0 ↔ Subsingleton M := by - let s := {k | lowerCentralSeries R L M k = ⊥} +theorem nilpotencyLength_eq_zero_iff [IsNilpotent L M] : + nilpotencyLength L M = 0 ↔ Subsingleton M := by + let s := {k | lowerCentralSeries ℤ L M k = ⊥} have hs : s.Nonempty := by - obtain ⟨k, hk⟩ := (by infer_instance : IsNilpotent R L M) + obtain ⟨k, hk⟩ := IsNilpotent.nilpotent ℤ L M exact ⟨k, hk⟩ change sInf s = 0 ↔ _ - rw [← LieSubmodule.subsingleton_iff R L M, ← subsingleton_iff_bot_eq_top, ← - lowerCentralSeries_zero, @eq_comm (LieSubmodule R L M)] + rw [← LieSubmodule.subsingleton_iff ℤ L M, ← subsingleton_iff_bot_eq_top, ← + lowerCentralSeries_zero, @eq_comm (LieSubmodule ℤ L M)] refine ⟨fun h => h ▸ Nat.sInf_mem hs, fun h => ?_⟩ rw [Nat.sInf_eq_zero] exact Or.inl h +section + +variable [LieModule R L M] + theorem nilpotencyLength_eq_succ_iff (k : ℕ) : - nilpotencyLength R L M = k + 1 ↔ + nilpotencyLength L M = k + 1 ↔ lowerCentralSeries R L M (k + 1) = ⊥ ∧ lowerCentralSeries R L M k ≠ ⊥ := by - let s := {k | lowerCentralSeries R L M k = ⊥} + have aux (k : ℕ) : lowerCentralSeries R L M k = ⊥ ↔ lowerCentralSeries ℤ L M k = ⊥ := by + simp [SetLike.ext'_iff, coe_lowerCentralSeries_eq_int R L M] + let s := {k | lowerCentralSeries ℤ L M k = ⊥} + rw [aux, ne_eq, aux] change sInf s = k + 1 ↔ k + 1 ∈ s ∧ k ∉ s have hs : ∀ k₁ k₂, k₁ ≤ k₂ → k₁ ∈ s → k₂ ∈ s := by - rintro k₁ k₂ h₁₂ (h₁ : lowerCentralSeries R L M k₁ = ⊥) - exact eq_bot_iff.mpr (h₁ ▸ antitone_lowerCentralSeries R L M h₁₂) + rintro k₁ k₂ h₁₂ (h₁ : lowerCentralSeries ℤ L M k₁ = ⊥) + exact eq_bot_iff.mpr (h₁ ▸ antitone_lowerCentralSeries ℤ L M h₁₂) exact Nat.sInf_upward_closed_eq_succ_iff hs k @[simp] theorem nilpotencyLength_eq_one_iff [Nontrivial M] : - nilpotencyLength R L M = 1 ↔ IsTrivial L M := by - rw [nilpotencyLength_eq_succ_iff, ← trivial_iff_lower_central_eq_bot] + nilpotencyLength L M = 1 ↔ IsTrivial L M := by + rw [nilpotencyLength_eq_succ_iff ℤ, ← trivial_iff_lower_central_eq_bot] simp -theorem isTrivial_of_nilpotencyLength_le_one [IsNilpotent R L M] (h : nilpotencyLength R L M ≤ 1) : +theorem isTrivial_of_nilpotencyLength_le_one [IsNilpotent L M] (h : nilpotencyLength L M ≤ 1) : IsTrivial L M := by nontriviality M cases' Nat.le_one_iff_eq_zero_or_eq_one.mp h with h h · rw [nilpotencyLength_eq_zero_iff] at h; infer_instance · rwa [nilpotencyLength_eq_one_iff] at h +end + /-- Given a non-trivial nilpotent Lie module `M` with lower central series `M = C₀ ≥ C₁ ≥ ⋯ ≥ Cₖ = ⊥`, this is the `k-1`th term in the lower central series (the last non-trivial term). For a trivial or non-nilpotent module, this is the bottom submodule, `⊥`. -/ noncomputable def lowerCentralSeriesLast : LieSubmodule R L M := - match nilpotencyLength R L M with + match nilpotencyLength L M with | 0 => ⊥ | k + 1 => lowerCentralSeries R L M k theorem lowerCentralSeriesLast_le_max_triv [LieModule R L M] : lowerCentralSeriesLast R L M ≤ maxTrivSubmodule R L M := by rw [lowerCentralSeriesLast] - cases' h : nilpotencyLength R L M with k + cases' h : nilpotencyLength L M with k · exact bot_le · rw [le_max_triv_iff_bracket_eq_bot] - rw [nilpotencyLength_eq_succ_iff, lowerCentralSeries_succ] at h + rw [nilpotencyLength_eq_succ_iff R, lowerCentralSeries_succ] at h exact h.1 -theorem nontrivial_lowerCentralSeriesLast [Nontrivial M] [IsNilpotent R L M] : +theorem nontrivial_lowerCentralSeriesLast [LieModule R L M] [Nontrivial M] [IsNilpotent L M] : Nontrivial (lowerCentralSeriesLast R L M) := by rw [LieSubmodule.nontrivial_iff_ne_bot, lowerCentralSeriesLast] - cases h : nilpotencyLength R L M + cases h : nilpotencyLength L M · rw [nilpotencyLength_eq_zero_iff, ← not_nontrivial_iff_subsingleton] at h contradiction - · rw [nilpotencyLength_eq_succ_iff] at h + · rw [nilpotencyLength_eq_succ_iff R] at h exact h.2 -theorem lowerCentralSeriesLast_le_of_not_isTrivial [IsNilpotent R L M] (h : ¬ IsTrivial L M) : +theorem lowerCentralSeriesLast_le_of_not_isTrivial [IsNilpotent L M] (h : ¬ IsTrivial L M) : lowerCentralSeriesLast R L M ≤ lowerCentralSeries R L M 1 := by rw [lowerCentralSeriesLast] - replace h : 1 < nilpotencyLength R L M := by + replace h : 1 < nilpotencyLength L M := by by_contra contra - have := isTrivial_of_nilpotencyLength_le_one R L M (not_lt.mp contra) + have := isTrivial_of_nilpotencyLength_le_one L M (not_lt.mp contra) contradiction - cases' hk : nilpotencyLength R L M with k <;> rw [hk] at h + cases' hk : nilpotencyLength L M with k <;> rw [hk] at h · contradiction · exact antitone_lowerCentralSeries _ _ _ (Nat.lt_succ.mp h) @@ -403,7 +463,7 @@ of `M` contains a non-zero element on which `L` acts trivially unless the entire Taking `M = L`, this provides a useful characterisation of Abelian-ness for nilpotent Lie algebras. -/ -lemma disjoint_lowerCentralSeries_maxTrivSubmodule_iff [IsNilpotent R L M] : +lemma disjoint_lowerCentralSeries_maxTrivSubmodule_iff [IsNilpotent L M] : Disjoint (lowerCentralSeries R L M 1) (maxTrivSubmodule R L M) ↔ IsTrivial L M := by refine ⟨fun h ↦ ?_, fun h ↦ by simp⟩ nontriviality M @@ -416,7 +476,7 @@ lemma disjoint_lowerCentralSeries_maxTrivSubmodule_iff [IsNilpotent R L M] : rw [h.eq_bot, le_bot_iff] at this exact this ▸ not_nontrivial _ -theorem nontrivial_max_triv_of_isNilpotent [Nontrivial M] [IsNilpotent R L M] : +theorem nontrivial_max_triv_of_isNilpotent [Nontrivial M] [IsNilpotent L M] : Nontrivial (maxTrivSubmodule R L M) := Set.nontrivial_mono (lowerCentralSeriesLast_le_max_triv R L M) (nontrivial_lowerCentralSeriesLast R L M) @@ -441,7 +501,8 @@ theorem coe_lcs_range_toEnd_eq (k : ℕ) : @[simp] theorem isNilpotent_range_toEnd_iff : - IsNilpotent R (toEnd R L M).range M ↔ IsNilpotent R L M := by + IsNilpotent (toEnd R L M).range M ↔ IsNilpotent L M := by + simp only [isNilpotent_iff R _ M] constructor <;> rintro ⟨k, hk⟩ <;> use k <;> rw [← LieSubmodule.toSubmodule_inj] at hk ⊢ <;> simpa using hk @@ -513,9 +574,10 @@ theorem gc_lcs_ucs (k : ℕ) : theorem ucs_eq_top_iff (k : ℕ) : N.ucs k = ⊤ ↔ LieModule.lowerCentralSeries R L M k ≤ N := by rw [eq_top_iff, ← lcs_le_iff]; rfl +variable (R) in theorem _root_.LieModule.isNilpotent_iff_exists_ucs_eq_top : - LieModule.IsNilpotent R L M ↔ ∃ k, (⊥ : LieSubmodule R L M).ucs k = ⊤ := by - rw [LieModule.isNilpotent_iff]; exact exists_congr fun k => by simp [ucs_eq_top_iff] + LieModule.IsNilpotent L M ↔ ∃ k, (⊥ : LieSubmodule R L M).ucs k = ⊤ := by + rw [LieModule.isNilpotent_iff R]; exact exists_congr fun k => by simp [ucs_eq_top_iff] theorem ucs_comap_incl (k : ℕ) : ((⊥ : LieSubmodule R L M).ucs k).comap N.incl = (⊥ : LieSubmodule R L N).ucs k := by @@ -524,8 +586,8 @@ theorem ucs_comap_incl (k : ℕ) : | succ k ih => simp [← ih] theorem isNilpotent_iff_exists_self_le_ucs : - LieModule.IsNilpotent R L N ↔ ∃ k, N ≤ (⊥ : LieSubmodule R L M).ucs k := by - simp_rw [LieModule.isNilpotent_iff_exists_ucs_eq_top, ← ucs_comap_incl, comap_incl_eq_top] + LieModule.IsNilpotent L N ↔ ∃ k, N ≤ (⊥ : LieSubmodule R L M).ucs k := by + simp_rw [LieModule.isNilpotent_iff_exists_ucs_eq_top R, ← ucs_comap_incl, comap_incl_eq_top] theorem ucs_bot_one : (⊥ : LieSubmodule R L M).ucs 1 = LieModule.maxTrivSubmodule R L M := by simp [LieSubmodule.normalizer_bot_eq_maxTrivSubmodule] @@ -568,14 +630,15 @@ theorem Function.Surjective.lieModule_lcs_map_eq (k : ℕ) : exact ⟨⁅y, n⁆, ⟨y, n, hn, rfl⟩, (hfg y n).symm⟩ include hf hg hfg in -theorem Function.Surjective.lieModuleIsNilpotent [IsNilpotent R L M] : IsNilpotent R L₂ M₂ := by - obtain ⟨k, hk⟩ := id (by infer_instance : IsNilpotent R L M) +theorem Function.Surjective.lieModuleIsNilpotent [IsNilpotent L M] : IsNilpotent L₂ M₂ := by + obtain ⟨k, hk⟩ := IsNilpotent.nilpotent R L M + rw [isNilpotent_iff R] use k rw [← LieSubmodule.toSubmodule_inj] at hk ⊢ simp [← hf.lieModule_lcs_map_eq hg hfg k, hk] theorem Equiv.lieModule_isNilpotent_iff (f : L ≃ₗ⁅R⁆ L₂) (g : M ≃ₗ[R] M₂) - (hfg : ∀ x m, ⁅f x, g m⁆ = g ⁅x, m⁆) : IsNilpotent R L M ↔ IsNilpotent R L₂ M₂ := by + (hfg : ∀ x m, ⁅f x, g m⁆ = g ⁅x, m⁆) : IsNilpotent L M ↔ IsNilpotent L₂ M₂ := by constructor <;> intro h · have hg : Surjective (g : M →ₗ[R] M₂) := g.surjective exact f.surjective.lieModuleIsNilpotent hg hfg @@ -586,38 +649,37 @@ theorem Equiv.lieModule_isNilpotent_iff (f : L ≃ₗ⁅R⁆ L₂) (g : M ≃ₗ @[simp] theorem LieModule.isNilpotent_of_top_iff : - IsNilpotent R (⊤ : LieSubalgebra R L) M ↔ IsNilpotent R L M := + IsNilpotent (⊤ : LieSubalgebra R L) M ↔ IsNilpotent L M := Equiv.lieModule_isNilpotent_iff LieSubalgebra.topEquiv (1 : M ≃ₗ[R] M) fun _ _ => rfl @[simp] lemma LieModule.isNilpotent_of_top_iff' : - IsNilpotent R L {x // x ∈ (⊤ : LieSubmodule R L M)} ↔ IsNilpotent R L M := + IsNilpotent L {x // x ∈ (⊤ : LieSubmodule R L M)} ↔ IsNilpotent L M := Equiv.lieModule_isNilpotent_iff 1 (LinearEquiv.ofTop ⊤ rfl) fun _ _ ↦ rfl end Morphisms end NilpotentModules -instance (priority := 100) LieAlgebra.isSolvable_of_isNilpotent (R : Type u) (L : Type v) - [CommRing R] [LieRing L] [LieAlgebra R L] [hL : LieModule.IsNilpotent R L L] : - LieAlgebra.IsSolvable R L := by - obtain ⟨k, h⟩ : ∃ k, LieModule.lowerCentralSeries R L L k = ⊥ := hL.nilpotent +instance (priority := 100) LieAlgebra.isSolvable_of_isNilpotent (L : Type v) + [LieRing L] [hL : LieModule.IsNilpotent L L] : + LieAlgebra.IsSolvable L := by + obtain ⟨k, h⟩ : ∃ k, LieModule.lowerCentralSeries ℤ L L k = ⊥ := hL.nilpotent_int use k; rw [← le_bot_iff] at h ⊢ - exact le_trans (LieModule.derivedSeries_le_lowerCentralSeries R L k) h + exact le_trans (LieModule.derivedSeries_le_lowerCentralSeries ℤ L k) h section NilpotentAlgebras variable (R : Type u) (L : Type v) (L' : Type w) variable [CommRing R] [LieRing L] [LieAlgebra R L] [LieRing L'] [LieAlgebra R L'] -/-- We say a Lie algebra is nilpotent when it is nilpotent as a Lie module over itself via the +/-- We say a Lie ring is nilpotent when it is nilpotent as a Lie module over itself via the adjoint representation. -/ -abbrev LieAlgebra.IsNilpotent (R : Type u) (L : Type v) [CommRing R] [LieRing L] [LieAlgebra R L] : - Prop := - LieModule.IsNilpotent R L L +abbrev LieRing.IsNilpotent (L : Type v) [LieRing L] : Prop := + LieModule.IsNilpotent L L -open LieAlgebra +open LieRing -theorem LieAlgebra.nilpotent_ad_of_nilpotent_algebra [IsNilpotent R L] : +theorem LieAlgebra.nilpotent_ad_of_nilpotent_algebra [IsNilpotent L] : ∃ k : ℕ, ∀ x : L, ad R L x ^ k = 0 := LieModule.exists_forall_pow_toEnd_eq_zero R L L @@ -667,14 +729,14 @@ theorem LieModule.coe_lowerCentralSeries_ideal_le {I : LieIdeal R L} (k : ℕ) : /-- A central extension of nilpotent Lie algebras is nilpotent. -/ theorem LieAlgebra.nilpotent_of_nilpotent_quotient {I : LieIdeal R L} (h₁ : I ≤ center R L) - (h₂ : IsNilpotent R (L ⧸ I)) : IsNilpotent R L := by - suffices LieModule.IsNilpotent R L (L ⧸ I) by + (h₂ : IsNilpotent (L ⧸ I)) : IsNilpotent L := by + suffices LieModule.IsNilpotent L (L ⧸ I) by exact LieModule.nilpotentOfNilpotentQuotient R L L h₁ this - obtain ⟨k, hk⟩ := h₂ - use k + simp only [LieRing.IsNilpotent, LieModule.isNilpotent_iff R] at h₂ ⊢ + peel h₂ with k hk simp [← LieSubmodule.toSubmodule_inj, coe_lowerCentralSeries_ideal_quot_eq, hk] -theorem LieAlgebra.non_trivial_center_of_isNilpotent [Nontrivial L] [IsNilpotent R L] : +theorem LieAlgebra.non_trivial_center_of_isNilpotent [Nontrivial L] [IsNilpotent L] : Nontrivial <| center R L := LieModule.nontrivial_max_triv_of_isNilpotent R L L @@ -695,29 +757,27 @@ theorem LieIdeal.lowerCentralSeries_map_eq (k : ℕ) {f : L →ₗ⁅R⁆ L'} (h | zero => simp only [LieModule.lowerCentralSeries_zero]; exact h' | succ k ih => simp only [LieModule.lowerCentralSeries_succ, LieIdeal.map_bracket_eq f h, ih, h'] -theorem Function.Injective.lieAlgebra_isNilpotent [h₁ : IsNilpotent R L'] {f : L →ₗ⁅R⁆ L'} - (h₂ : Function.Injective f) : IsNilpotent R L := - { nilpotent := by - obtain ⟨k, hk⟩ := id h₁ - use k - apply LieIdeal.bot_of_map_eq_bot h₂; rw [eq_bot_iff, ← hk] - apply LieIdeal.map_lowerCentralSeries_le } - -theorem Function.Surjective.lieAlgebra_isNilpotent [h₁ : IsNilpotent R L] {f : L →ₗ⁅R⁆ L'} - (h₂ : Function.Surjective f) : IsNilpotent R L' := - { nilpotent := by - obtain ⟨k, hk⟩ := id h₁ - use k - rw [← LieIdeal.lowerCentralSeries_map_eq k h₂, hk] - simp only [LieIdeal.map_eq_bot_iff, bot_le] } +theorem Function.Injective.lieAlgebra_isNilpotent [h₁ : IsNilpotent L'] {f : L →ₗ⁅R⁆ L'} + (h₂ : Function.Injective f) : IsNilpotent L := by + rw [LieRing.IsNilpotent, LieModule.isNilpotent_iff R] at h₁ ⊢ + peel h₁ with k hk + apply LieIdeal.bot_of_map_eq_bot h₂; rw [eq_bot_iff, ← hk] + apply LieIdeal.map_lowerCentralSeries_le + +theorem Function.Surjective.lieAlgebra_isNilpotent [h₁ : IsNilpotent L] {f : L →ₗ⁅R⁆ L'} + (h₂ : Function.Surjective f) : IsNilpotent L' := by + rw [LieRing.IsNilpotent, LieModule.isNilpotent_iff R] at h₁ ⊢ + peel h₁ with k hk + rw [← LieIdeal.lowerCentralSeries_map_eq k h₂, hk] + simp only [LieIdeal.map_eq_bot_iff, bot_le] theorem LieEquiv.nilpotent_iff_equiv_nilpotent (e : L ≃ₗ⁅R⁆ L') : - IsNilpotent R L ↔ IsNilpotent R L' := by + IsNilpotent L ↔ IsNilpotent L' := by constructor <;> intro h · exact e.symm.injective.lieAlgebra_isNilpotent · exact e.injective.lieAlgebra_isNilpotent -theorem LieHom.isNilpotent_range [IsNilpotent R L] (f : L →ₗ⁅R⁆ L') : IsNilpotent R f.range := +theorem LieHom.isNilpotent_range [IsNilpotent L] (f : L →ₗ⁅R⁆ L') : IsNilpotent f.range := f.surjective_rangeRestrict.lieAlgebra_isNilpotent /-- Note that this result is not quite a special case of @@ -725,7 +785,7 @@ theorem LieHom.isNilpotent_range [IsNilpotent R L] (f : L →ₗ⁅R⁆ L') : Is `(ad R L).range`-module `L`, whereas this result concerns nilpotency of the `(ad R L).range`-module `(ad R L).range`. -/ @[simp] -theorem LieAlgebra.isNilpotent_range_ad_iff : IsNilpotent R (ad R L).range ↔ IsNilpotent R L := by +theorem LieAlgebra.isNilpotent_range_ad_iff : IsNilpotent (ad R L).range ↔ IsNilpotent L := by refine ⟨fun h => ?_, ?_⟩ · have : (ad R L).ker = center R L := by simp exact @@ -734,7 +794,7 @@ theorem LieAlgebra.isNilpotent_range_ad_iff : IsNilpotent R (ad R L).range ↔ I · intro h exact (ad R L).isNilpotent_range -instance [h : LieAlgebra.IsNilpotent R L] : LieAlgebra.IsNilpotent R (⊤ : LieSubalgebra R L) := +instance [h : LieRing.IsNilpotent L] : LieRing.IsNilpotent (⊤ : LieSubalgebra R L) := LieSubalgebra.topEquiv.nilpotent_iff_equiv_nilpotent.mpr h end NilpotentAlgebras @@ -829,9 +889,10 @@ lemma LieSubmodule.lowerCentralSeries_tensor_eq_baseChange (k : ℕ) : | zero => simp | succ k ih => simp only [lowerCentralSeries_succ, ih, ← baseChange_top, lie_baseChange] -instance LieModule.instIsNilpotentTensor [IsNilpotent R L M] : - IsNilpotent A (A ⊗[R] L) (A ⊗[R] M) := by - obtain ⟨k, hk⟩ := inferInstanceAs (IsNilpotent R L M) +instance LieModule.instIsNilpotentTensor [IsNilpotent L M] : + IsNilpotent (A ⊗[R] L) (A ⊗[R] M) := by + obtain ⟨k, hk⟩ := IsNilpotent.nilpotent R L M + rw [isNilpotent_iff A] exact ⟨k, by simp [hk]⟩ end ExtendScalars diff --git a/Mathlib/Algebra/Lie/Semisimple/Basic.lean b/Mathlib/Algebra/Lie/Semisimple/Basic.lean index 16117bb7c15bf..0697232c790f3 100644 --- a/Mathlib/Algebra/Lie/Semisimple/Basic.lean +++ b/Mathlib/Algebra/Lie/Semisimple/Basic.lean @@ -48,7 +48,7 @@ variable (R L : Type*) [CommRing R] [LieRing L] [LieAlgebra R L] variable {R L} in theorem HasTrivialRadical.eq_bot_of_isSolvable [HasTrivialRadical R L] - (I : LieIdeal R L) [hI : IsSolvable R I] : I = ⊥ := + (I : LieIdeal R L) [hI : IsSolvable I] : I = ⊥ := sSup_eq_bot.mp radical_eq_bot _ hI @[simp] @@ -56,19 +56,19 @@ theorem HasTrivialRadical.center_eq_bot [HasTrivialRadical R L] : center R L = HasTrivialRadical.eq_bot_of_isSolvable _ variable {R L} in -theorem hasTrivialRadical_of_no_solvable_ideals (h : ∀ I : LieIdeal R L, IsSolvable R I → I = ⊥) : +theorem hasTrivialRadical_of_no_solvable_ideals (h : ∀ I : LieIdeal R L, IsSolvable I → I = ⊥) : HasTrivialRadical R L := ⟨sSup_eq_bot.mpr h⟩ theorem hasTrivialRadical_iff_no_solvable_ideals : - HasTrivialRadical R L ↔ ∀ I : LieIdeal R L, IsSolvable R I → I = ⊥ := + HasTrivialRadical R L ↔ ∀ I : LieIdeal R L, IsSolvable I → I = ⊥ := ⟨@HasTrivialRadical.eq_bot_of_isSolvable _ _ _ _ _, hasTrivialRadical_of_no_solvable_ideals⟩ theorem hasTrivialRadical_iff_no_abelian_ideals : HasTrivialRadical R L ↔ ∀ I : LieIdeal R L, IsLieAbelian I → I = ⊥ := by rw [hasTrivialRadical_iff_no_solvable_ideals] constructor <;> intro h₁ I h₂ - · exact h₁ _ <| LieAlgebra.ofAbelianIsSolvable R I + · exact h₁ _ <| LieAlgebra.ofAbelianIsSolvable I · rw [← abelian_of_solvable_ideal_eq_bot_iff] exact h₁ _ <| abelian_derivedAbelianOfIdeal I @@ -318,7 +318,7 @@ to be reductive. Note that there is absolutely [no agreement](https://mathoverflow.net/questions/284713/) on what the label 'reductive' should mean when the coefficients are not a field of characteristic zero. -/ theorem abelian_radical_iff_solvable_is_abelian [IsNoetherian R L] : - IsLieAbelian (radical R L) ↔ ∀ I : LieIdeal R L, IsSolvable R I → IsLieAbelian I := by + IsLieAbelian (radical R L) ↔ ∀ I : LieIdeal R L, IsSolvable I → IsLieAbelian I := by constructor · rintro h₁ I h₂ rw [LieIdeal.solvable_iff_le_radical] at h₂ diff --git a/Mathlib/Algebra/Lie/Solvable.lean b/Mathlib/Algebra/Lie/Solvable.lean index a17e1accef60b..7e9c2d579d51f 100644 --- a/Mathlib/Algebra/Lie/Solvable.lean +++ b/Mathlib/Algebra/Lie/Solvable.lean @@ -190,23 +190,71 @@ theorem derivedSeries_eq_top (n : ℕ) (h : derivedSeries R L 1 = ⊤) : · rfl · rwa [derivedSeries_succ_eq_top_iff] +private theorem coe_derivedSeries_eq_int_aux (R₁ R₂ L : Type*) [CommRing R₁] [CommRing R₂] + [LieRing L] [LieAlgebra R₁ L] [LieAlgebra R₂ L] (k : ℕ) + (ih : ∀ (x : L), x ∈ derivedSeriesOfIdeal R₁ L k ⊤ ↔ x ∈ derivedSeriesOfIdeal R₂ L k ⊤) : + let I := derivedSeriesOfIdeal R₂ L k ⊤; let S : Set L := {⁅a, b⁆ | (a ∈ I) (b ∈ I)} + (Submodule.span R₁ S : Set L) ≤ (Submodule.span R₂ S : Set L) := by + intro I S x hx + simp only [SetLike.mem_coe] at hx ⊢ + induction hx using Submodule.closure_induction with + | zero => exact Submodule.zero_mem _ + | add y z hy₁ hz₁ hy₂ hz₂ => exact Submodule.add_mem _ hy₂ hz₂ + | smul_mem c y hy => + obtain ⟨a, ha, b, hb, rfl⟩ := hy + rw [← smul_lie] + refine Submodule.subset_span ⟨c • a, ?_, b, hb, rfl⟩ + rw [← ih] at ha ⊢ + exact Submodule.smul_mem _ _ ha + +theorem coe_derivedSeries_eq_int (k : ℕ) : + (derivedSeries R L k : Set L) = (derivedSeries ℤ L k : Set L) := by + show ((derivedSeries R L k).toSubmodule : Set L) = ((derivedSeries ℤ L k).toSubmodule : Set L) + rw [derivedSeries_def, derivedSeries_def] + induction k with + | zero => rfl + | succ k ih => + rw [derivedSeriesOfIdeal_succ, derivedSeriesOfIdeal_succ] + rw [LieSubmodule.lieIdeal_oper_eq_linear_span', LieSubmodule.lieIdeal_oper_eq_linear_span'] + rw [Set.ext_iff] at ih + simp only [SetLike.mem_coe, LieSubmodule.mem_toSubmodule] at ih + simp only [Subtype.exists, exists_prop, ih] + apply le_antisymm + · exact coe_derivedSeries_eq_int_aux _ _ L k ih + · simp only [← ih] + apply coe_derivedSeries_eq_int_aux _ _ L k + simp [ih] + end LieIdeal namespace LieAlgebra /-- A Lie algebra is solvable if its derived series reaches 0 (in a finite number of steps). -/ +@[mk_iff isSolvable_iff_int] class IsSolvable : Prop where - solvable : ∃ k, derivedSeries R L k = ⊥ + mk_int :: + solvable_int : ∃ k, derivedSeries ℤ L k = ⊥ -instance isSolvableBot : IsSolvable R (⊥ : LieIdeal R L) := +instance isSolvableBot : IsSolvable (⊥ : LieIdeal R L) := ⟨⟨0, Subsingleton.elim _ ⊥⟩⟩ -instance isSolvableAdd {I J : LieIdeal R L} [hI : IsSolvable R I] [hJ : IsSolvable R J] : - IsSolvable R (I + J) := by - obtain ⟨k, hk⟩ := id hI; obtain ⟨l, hl⟩ := id hJ - exact ⟨⟨k + l, LieIdeal.derivedSeries_add_eq_bot hk hl⟩⟩ +lemma isSolvable_iff : IsSolvable L ↔ ∃ k, derivedSeries R L k = ⊥ := by + simp [isSolvable_iff_int, SetLike.ext'_iff, LieIdeal.coe_derivedSeries_eq_int] + +lemma IsSolvable.solvable [IsSolvable L] : ∃ k, derivedSeries R L k = ⊥ := + (isSolvable_iff R L).mp ‹_› + +variable {R L} in +lemma IsSolvable.mk {k : ℕ} (h : derivedSeries R L k = ⊥) : IsSolvable L := + (isSolvable_iff R L).mpr ⟨k, h⟩ + +instance isSolvableAdd {I J : LieIdeal R L} [IsSolvable I] [IsSolvable J] : + IsSolvable (I + J) := by + obtain ⟨k, hk⟩ := IsSolvable.solvable R I + obtain ⟨l, hl⟩ := IsSolvable.solvable R J + exact IsSolvable.mk (LieIdeal.derivedSeries_add_eq_bot hk hl) -theorem derivedSeries_lt_top_of_solvable [IsSolvable R L] [Nontrivial L] : +theorem derivedSeries_lt_top_of_solvable [IsSolvable L] [Nontrivial L] : derivedSeries R L 1 < ⊤ := by obtain ⟨n, hn⟩ := IsSolvable.solvable (R := R) (L := L) rw [lt_top_iff_ne_top] @@ -222,76 +270,78 @@ namespace Function open LieAlgebra -theorem Injective.lieAlgebra_isSolvable [h₁ : IsSolvable R L] (h₂ : Injective f) : - IsSolvable R L' := by - obtain ⟨k, hk⟩ := id h₁ - use k - apply LieIdeal.bot_of_map_eq_bot h₂; rw [eq_bot_iff, ← hk] +theorem Injective.lieAlgebra_isSolvable [hL : IsSolvable L] (h : Injective f) : + IsSolvable L' := by + rw [isSolvable_iff R] at hL ⊢ + apply hL.imp + intro k hk + apply LieIdeal.bot_of_map_eq_bot h; rw [eq_bot_iff, ← hk] apply LieIdeal.derivedSeries_map_le -instance (A : LieIdeal R L) [IsSolvable R L] : IsSolvable R A := +instance (A : LieIdeal R L) [IsSolvable L] : IsSolvable A := A.incl_injective.lieAlgebra_isSolvable -theorem Surjective.lieAlgebra_isSolvable [h₁ : IsSolvable R L'] (h₂ : Surjective f) : - IsSolvable R L := by - obtain ⟨k, hk⟩ := id h₁ - use k - rw [← LieIdeal.derivedSeries_map_eq k h₂, hk] +theorem Surjective.lieAlgebra_isSolvable [hL' : IsSolvable L'] (h : Surjective f) : + IsSolvable L := by + rw [isSolvable_iff R] at hL' ⊢ + apply hL'.imp + intro k hk + rw [← LieIdeal.derivedSeries_map_eq k h, hk] simp only [LieIdeal.map_eq_bot_iff, bot_le] end Function -instance LieHom.isSolvable_range (f : L' →ₗ⁅R⁆ L) [LieAlgebra.IsSolvable R L'] : - LieAlgebra.IsSolvable R f.range := +instance LieHom.isSolvable_range (f : L' →ₗ⁅R⁆ L) [LieAlgebra.IsSolvable L'] : + LieAlgebra.IsSolvable f.range := f.surjective_rangeRestrict.lieAlgebra_isSolvable namespace LieAlgebra -theorem solvable_iff_equiv_solvable (e : L' ≃ₗ⁅R⁆ L) : IsSolvable R L' ↔ IsSolvable R L := by +theorem solvable_iff_equiv_solvable (e : L' ≃ₗ⁅R⁆ L) : IsSolvable L' ↔ IsSolvable L := by constructor <;> intro h · exact e.symm.injective.lieAlgebra_isSolvable · exact e.injective.lieAlgebra_isSolvable -theorem le_solvable_ideal_solvable {I J : LieIdeal R L} (h₁ : I ≤ J) (_ : IsSolvable R J) : - IsSolvable R I := +theorem le_solvable_ideal_solvable {I J : LieIdeal R L} (h₁ : I ≤ J) (_ : IsSolvable J) : + IsSolvable I := (LieIdeal.inclusion_injective h₁).lieAlgebra_isSolvable variable (R L) -instance (priority := 100) ofAbelianIsSolvable [IsLieAbelian L] : IsSolvable R L := by +instance (priority := 100) ofAbelianIsSolvable [IsLieAbelian L] : IsSolvable L := by use 1 rw [← abelian_iff_derived_one_eq_bot, lie_abelian_iff_equiv_lie_abelian LieIdeal.topEquiv] infer_instance /-- The (solvable) radical of Lie algebra is the `sSup` of all solvable ideals. -/ def radical := - sSup { I : LieIdeal R L | IsSolvable R I } + sSup { I : LieIdeal R L | IsSolvable I } /-- The radical of a Noetherian Lie algebra is solvable. -/ -instance radicalIsSolvable [IsNoetherian R L] : IsSolvable R (radical R L) := by +instance radicalIsSolvable [IsNoetherian R L] : IsSolvable (radical R L) := by have hwf := LieSubmodule.wellFoundedGT_of_noetherian R L L rw [← CompleteLattice.isSupClosedCompact_iff_wellFoundedGT] at hwf - refine hwf { I : LieIdeal R L | IsSolvable R I } ⟨⊥, ?_⟩ fun I hI J hJ => ?_ + refine hwf { I : LieIdeal R L | IsSolvable I } ⟨⊥, ?_⟩ fun I hI J hJ => ?_ · exact LieAlgebra.isSolvableBot R L · rw [Set.mem_setOf_eq] at hI hJ ⊢ apply LieAlgebra.isSolvableAdd R L /-- The `→` direction of this lemma is actually true without the `IsNoetherian` assumption. -/ theorem LieIdeal.solvable_iff_le_radical [IsNoetherian R L] (I : LieIdeal R L) : - IsSolvable R I ↔ I ≤ radical R L := + IsSolvable I ↔ I ≤ radical R L := ⟨fun h => le_sSup h, fun h => le_solvable_ideal_solvable h inferInstance⟩ theorem center_le_radical : center R L ≤ radical R L := - have h : IsSolvable R (center R L) := inferInstance + have h : IsSolvable (center R L) := inferInstance le_sSup h -instance [IsSolvable R L] : IsSolvable R (⊤ : LieSubalgebra R L) := by +instance [IsSolvable L] : IsSolvable (⊤ : LieSubalgebra R L) := by rwa [solvable_iff_equiv_solvable LieSubalgebra.topEquiv] -@[simp] lemma radical_eq_top_of_isSolvable [IsSolvable R L] : +@[simp] lemma radical_eq_top_of_isSolvable [IsSolvable L] : radical R L = ⊤ := by rw [eq_top_iff] - have h : IsSolvable R (⊤ : LieSubalgebra R L) := inferInstance + have h : IsSolvable (⊤ : LieSubalgebra R L) := inferInstance exact le_sSup h /-- Given a solvable Lie ideal `I` with derived series `I = D₀ ≥ D₁ ≥ ⋯ ≥ Dₖ = ⊥`, this is the @@ -348,17 +398,17 @@ theorem abelian_derivedAbelianOfIdeal (I : LieIdeal R L) : · dsimp; infer_instance · rw [derivedSeries_of_derivedLength_succ] at h; exact h.1 -theorem derivedLength_zero (I : LieIdeal R L) [hI : IsSolvable R I] : +theorem derivedLength_zero (I : LieIdeal R L) [IsSolvable I] : derivedLengthOfIdeal R L I = 0 ↔ I = ⊥ := by let s := { k | derivedSeriesOfIdeal R L k I = ⊥ } change sInf s = 0 ↔ _ have hne : s ≠ ∅ := by - obtain ⟨k, hk⟩ := id hI + obtain ⟨k, hk⟩ := IsSolvable.solvable R I refine Set.Nonempty.ne_empty ⟨k, ?_⟩ rw [derivedSeries_def, LieIdeal.derivedSeries_eq_bot_iff] at hk; exact hk simp [s, hne] -theorem abelian_of_solvable_ideal_eq_bot_iff (I : LieIdeal R L) [h : IsSolvable R I] : +theorem abelian_of_solvable_ideal_eq_bot_iff (I : LieIdeal R L) [h : IsSolvable I] : derivedAbelianOfIdeal I = ⊥ ↔ I = ⊥ := by dsimp only [derivedAbelianOfIdeal] split -- Porting note: Original tactic was `cases' h : derivedAbelianOfIdeal R L I with k` diff --git a/Mathlib/Algebra/Lie/TraceForm.lean b/Mathlib/Algebra/Lie/TraceForm.lean index e5914475629d3..80c810d875cb6 100644 --- a/Mathlib/Algebra/Lie/TraceForm.lean +++ b/Mathlib/Algebra/Lie/TraceForm.lean @@ -94,7 +94,7 @@ lemma traceForm_lieInvariant : (traceForm R L M).lieInvariant L := by rw [LieHom.lie_apply, LinearMap.sub_apply, Module.Dual.lie_apply, LinearMap.zero_apply, LinearMap.zero_apply, traceForm_apply_lie_apply', sub_self] -@[simp] lemma traceForm_eq_zero_of_isNilpotent [IsReduced R] [IsNilpotent R L M] : +@[simp] lemma traceForm_eq_zero_of_isNilpotent [IsReduced R] [IsNilpotent L M] : traceForm R L M = 0 := by ext x y simp only [traceForm_apply_apply, LinearMap.zero_apply, ← isNilpotent_iff_eq_zero] @@ -104,7 +104,7 @@ lemma traceForm_lieInvariant : (traceForm R L M).lieInvariant L := by @[simp] lemma traceForm_genWeightSpace_eq [Module.Free R M] [IsDomain R] [IsPrincipalIdealRing R] - [LieAlgebra.IsNilpotent R L] [IsNoetherian R M] [LinearWeights R L M] (χ : L → R) (x y : L) : + [LieRing.IsNilpotent L] [IsNoetherian R M] [LinearWeights R L M] (χ : L → R) (x y : L) : traceForm R L (genWeightSpace M χ) x y = finrank R (genWeightSpace M χ) • (χ x * χ y) := by set d := finrank R (genWeightSpace M χ) have h₁ : χ y • d • χ x - χ y • χ x • (d : R) = 0 := by simp [mul_comm (χ x)] @@ -160,7 +160,7 @@ lemma traceForm_apply_eq_zero_of_mem_lcs_of_mem_center {x y : L} /-- Given a bilinear form `B` on a representation `M` of a nilpotent Lie algebra `L`, if `B` is invariant (in the sense that the action of `L` is skew-adjoint wrt `B`) then components of the Fitting decomposition of `M` are orthogonal wrt `B`. -/ -lemma eq_zero_of_mem_genWeightSpace_mem_posFitting [LieAlgebra.IsNilpotent R L] +lemma eq_zero_of_mem_genWeightSpace_mem_posFitting [LieRing.IsNilpotent L] {B : LinearMap.BilinForm R M} (hB : ∀ (x : L) (m n : M), B ⁅x, m⁆ n = - B m ⁅x, n⁆) {m₀ m₁ : M} (hm₀ : m₀ ∈ genWeightSpace M (0 : L → R)) (hm₁ : m₁ ∈ posFittingComp R L M) : B m₀ m₁ = 0 := by @@ -211,7 +211,7 @@ lemma traceForm_lieSubalgebra_mk_right (L' : LieSubalgebra R L) {x : L'} {y : L} open TensorProduct -variable [LieAlgebra.IsNilpotent R L] [IsDomain R] [IsPrincipalIdealRing R] +variable [LieRing.IsNilpotent L] [IsDomain R] [IsPrincipalIdealRing R] lemma traceForm_eq_sum_genWeightSpaceOf [NoZeroSMulDivisors R M] [IsNoetherian R M] [IsTriangularizable R L M] (z : L) : @@ -344,7 +344,7 @@ lemma killingForm_apply_apply (x y : L) : killingForm R L x y = trace R L (ad R LieModule.traceForm_apply_apply R L L x y lemma killingForm_eq_zero_of_mem_zeroRoot_mem_posFitting - (H : LieSubalgebra R L) [LieAlgebra.IsNilpotent R H] + (H : LieSubalgebra R L) [LieRing.IsNilpotent H] {x₀ x₁ : L} (hx₀ : x₀ ∈ LieAlgebra.zeroRootSubalgebra R L H) (hx₁ : x₁ ∈ LieModule.posFittingComp R H L) : @@ -399,7 +399,7 @@ open Submodule (span subset_span) namespace LieModule variable [Field K] [LieAlgebra K L] [Module K M] [LieModule K L M] [FiniteDimensional K M] -variable [LieAlgebra.IsNilpotent K L] [LinearWeights K L M] [IsTriangularizable K L M] +variable [LieRing.IsNilpotent L] [LinearWeights K L M] [IsTriangularizable K L M] lemma traceForm_eq_sum_finrank_nsmul_mul (x y : L) : traceForm K L M x y = ∑ χ : Weight K L M, finrank K (genWeightSpace M χ) • (χ x * χ y) := by diff --git a/Mathlib/Algebra/Lie/Weights/Basic.lean b/Mathlib/Algebra/Lie/Weights/Basic.lean index 7ecb88285cfc7..aaee9a5ecd8b8 100644 --- a/Mathlib/Algebra/Lie/Weights/Basic.lean +++ b/Mathlib/Algebra/Lie/Weights/Basic.lean @@ -156,7 +156,7 @@ variable (M) `genWeightSpaceOf M χ x` is the maximal generalized `χ`-eigenspace of the action of `x` on `M`. It is a Lie submodule because `L` is nilpotent. -/ -def genWeightSpaceOf [LieAlgebra.IsNilpotent R L] (χ : R) (x : L) : LieSubmodule R L M := +def genWeightSpaceOf [LieRing.IsNilpotent L] (χ : R) (x : L) : LieSubmodule R L M := { 𝕎(M, χ, x) with lie_mem := by intro y m hm @@ -168,7 +168,7 @@ def genWeightSpaceOf [LieAlgebra.IsNilpotent R L] (χ : R) (x : L) : LieSubmodul end notation_genWeightSpaceOf variable (M) -variable [LieAlgebra.IsNilpotent R L] +variable [LieRing.IsNilpotent L] theorem mem_genWeightSpaceOf (χ : R) (x : L) (m : M) : m ∈ genWeightSpaceOf M χ x ↔ ∃ k : ℕ, ((toEnd R L M x - χ • ↑1) ^ k) m = 0 := by @@ -228,7 +228,7 @@ variable {M} @[ext] lemma ext {χ₁ χ₂ : Weight R L M} (h : ∀ x, χ₁ x = χ₂ x) : χ₁ = χ₂ := by cases' χ₁ with f₁ _; cases' χ₂ with f₂ _; aesop -lemma ext_iff' {χ₁ χ₂ : Weight R L M} : (χ₁ : L → R) = χ₂ ↔ χ₁ = χ₂ := by aesop +lemma ext_iff' {χ₁ χ₂ : Weight R L M} : (χ₁ : L → R) = χ₂ ↔ χ₁ = χ₂ := by simp lemma exists_ne_zero (χ : Weight R L M) : ∃ x ∈ genWeightSpace M χ, x ≠ 0 := by @@ -298,7 +298,7 @@ end Weight /-- See also the more useful form `LieModule.zero_genWeightSpace_eq_top_of_nilpotent`. -/ @[simp] -theorem zero_genWeightSpace_eq_top_of_nilpotent' [IsNilpotent R L M] : +theorem zero_genWeightSpace_eq_top_of_nilpotent' [IsNilpotent L M] : genWeightSpace M (0 : L → R) = ⊤ := by ext simp [genWeightSpace, genWeightSpaceOf] @@ -311,7 +311,7 @@ theorem coe_genWeightSpace_of_top (χ : L → R) : simp @[simp] -theorem zero_genWeightSpace_eq_top_of_nilpotent [IsNilpotent R L M] : +theorem zero_genWeightSpace_eq_top_of_nilpotent [IsNilpotent L M] : genWeightSpace M (0 : (⊤ : LieSubalgebra R L) → R) = ⊤ := by ext m simp only [mem_genWeightSpace, Pi.zero_apply, zero_smul, sub_zero, Subtype.forall, @@ -356,7 +356,7 @@ theorem isNilpotent_toEnd_genWeightSpace_zero [IsNoetherian R M] (x : L) : /-- By Engel's theorem, the zero weight space of a Noetherian Lie module is nilpotent. -/ instance [IsNoetherian R M] : - IsNilpotent R L (genWeightSpace M (0 : L → R)) := + IsNilpotent L (genWeightSpace M (0 : L → R)) := isNilpotent_iff_forall'.mpr <| isNilpotent_toEnd_genWeightSpace_zero M variable (R L) @@ -437,7 +437,7 @@ lemma mem_posFittingCompOf (x : L) (m : M) : exact LieSubmodule.lie_mem_lie (LieSubmodule.mem_top x) ih @[simp] lemma posFittingCompOf_eq_bot_of_isNilpotent - [IsNilpotent R L M] (x : L) : + [IsNilpotent L M] (x : L) : posFittingCompOf R M x = ⊥ := by simp_rw [eq_bot_iff, ← iInf_lowerCentralSeries_eq_bot_of_isNilpotent, le_iInf_iff, posFittingCompOf_le_lowerCentralSeries, forall_const] @@ -470,7 +470,7 @@ lemma posFittingComp_le_iInf_lowerCentralSeries : ⨅ k, lowerCentralSeries R L M k = posFittingComp R L M := by refine le_antisymm ?_ (posFittingComp_le_iInf_lowerCentralSeries R L M) apply iInf_lcs_le_of_isNilpotent_quot - rw [LieModule.isNilpotent_iff_forall'] + rw [LieModule.isNilpotent_iff_forall' (R := R)] intro x obtain ⟨k, hk⟩ := Filter.eventually_atTop.mp (toEnd R L M x).eventually_iInf_range_pow_eq use k @@ -488,7 +488,7 @@ lemma posFittingComp_le_iInf_lowerCentralSeries : simpa using this @[simp] lemma posFittingComp_eq_bot_of_isNilpotent - [IsNilpotent R L M] : + [IsNilpotent L M] : posFittingComp R L M = ⊥ := by simp [posFittingComp] @@ -603,9 +603,9 @@ private lemma isCompl_genWeightSpace_zero_posFittingComp_aux set M₁ := posFittingComp R L M rcases forall_or_exists_not (fun (x : L) ↦ genWeightSpaceOf M (0 : R) x = ⊤) with h | ⟨x, hx : genWeightSpaceOf M (0 : R) x ≠ ⊤⟩ - · suffices IsNilpotent R L M by simp [M₀, M₁, isCompl_top_bot] + · suffices IsNilpotent L M by simp [M₀, M₁, isCompl_top_bot] replace h : M₀ = ⊤ := by simpa [M₀, genWeightSpace] - rw [← LieModule.isNilpotent_of_top_iff', ← h] + rw [← LieModule.isNilpotent_of_top_iff' (R := R), ← h] infer_instance · set M₀ₓ := genWeightSpaceOf M (0 : R) x set M₁ₓ := posFittingCompOf R M x @@ -750,8 +750,7 @@ section field open Module variable (K) -variable [Field K] [LieAlgebra K L] [Module K M] [LieModule K L M] [LieAlgebra.IsNilpotent K L] - [FiniteDimensional K M] +variable [Field K] [LieAlgebra K L] [Module K M] [LieModule K L M] [FiniteDimensional K M] instance instIsTriangularizableOfIsAlgClosed [IsAlgClosed K] : IsTriangularizable K L M := ⟨fun _ ↦ Module.End.iSup_maxGenEigenspace_eq_top _⟩ diff --git a/Mathlib/Algebra/Lie/Weights/Cartan.lean b/Mathlib/Algebra/Lie/Weights/Cartan.lean index 17a743e544aa1..d979a55e0b938 100644 --- a/Mathlib/Algebra/Lie/Weights/Cartan.lean +++ b/Mathlib/Algebra/Lie/Weights/Cartan.lean @@ -31,7 +31,7 @@ suppress_compilation open Set variable {R L : Type*} [CommRing R] [LieRing L] [LieAlgebra R L] - (H : LieSubalgebra R L) [LieAlgebra.IsNilpotent R H] + (H : LieSubalgebra R L) [LieRing.IsNilpotent H] {M : Type*} [AddCommGroup M] [Module R M] [LieRingModule L M] [LieModule R L M] namespace LieAlgebra @@ -44,7 +44,7 @@ space of `L` regarded as a module of `H` via the adjoint action. -/ abbrev rootSpace (χ : H → R) : LieSubmodule R H L := genWeightSpace L χ -theorem zero_rootSpace_eq_top_of_nilpotent [IsNilpotent R L] : +theorem zero_rootSpace_eq_top_of_nilpotent [LieRing.IsNilpotent L] : rootSpace (⊤ : LieSubalgebra R L) 0 = ⊤ := zero_genWeightSpace_eq_top_of_nilpotent L @@ -170,7 +170,7 @@ theorem toLieSubmodule_le_rootSpace_zero : H.toLieSubmodule ≤ rootSpace H 0 := simp only [LieSubalgebra.mem_toLieSubmodule] at hx simp only [mem_genWeightSpace, Pi.zero_apply, sub_zero, zero_smul] intro y - obtain ⟨k, hk⟩ := (inferInstance : IsNilpotent R H) + obtain ⟨k, hk⟩ := IsNilpotent.nilpotent R H H use k let f : Module.End R H := toEnd R H H y let g : Module.End R L := toEnd R H L y diff --git a/Mathlib/Algebra/Lie/Weights/Chain.lean b/Mathlib/Algebra/Lie/Weights/Chain.lean index b2ba09bb4e2bd..14ab6cec6e4a9 100644 --- a/Mathlib/Algebra/Lie/Weights/Chain.lean +++ b/Mathlib/Algebra/Lie/Weights/Chain.lean @@ -50,7 +50,7 @@ namespace LieModule section IsNilpotent -variable [LieAlgebra.IsNilpotent R L] (χ₁ χ₂ : L → R) (p q : ℤ) +variable [LieRing.IsNilpotent L] (χ₁ χ₂ : L → R) (p q : ℤ) section @@ -121,7 +121,7 @@ open LieAlgebra variable {H : LieSubalgebra R L} (α χ : H → R) (p q : ℤ) -lemma lie_mem_genWeightSpaceChain_of_genWeightSpace_eq_bot_right [LieAlgebra.IsNilpotent R H] +lemma lie_mem_genWeightSpaceChain_of_genWeightSpace_eq_bot_right [LieRing.IsNilpotent H] (hq : genWeightSpace M (q • α + χ) = ⊥) {x : L} (hx : x ∈ rootSpace H α) {y : M} (hy : y ∈ genWeightSpaceChain M α χ p q) : @@ -144,7 +144,7 @@ lemma lie_mem_genWeightSpaceChain_of_genWeightSpace_eq_bot_right [LieAlgebra.IsN | h0 => simp | hadd _ _ _ _ hz₁ hz₂ => rw [lie_add]; exact add_mem hz₁ hz₂ -lemma lie_mem_genWeightSpaceChain_of_genWeightSpace_eq_bot_left [LieAlgebra.IsNilpotent R H] +lemma lie_mem_genWeightSpaceChain_of_genWeightSpace_eq_bot_left [LieRing.IsNilpotent H] (hp : genWeightSpace M (p • α + χ) = ⊥) {x : L} (hx : x ∈ rootSpace H (-α)) {y : M} (hy : y ∈ genWeightSpaceChain M α χ p q) : @@ -234,7 +234,7 @@ end LieSubalgebra section variable {M} -variable [LieAlgebra.IsNilpotent R L] +variable [LieRing.IsNilpotent L] variable [NoZeroSMulDivisors ℤ R] [NoZeroSMulDivisors R M] [IsNoetherian R M] variable (α : L → R) (β : Weight R L M) diff --git a/Mathlib/Algebra/Lie/Weights/Linear.lean b/Mathlib/Algebra/Lie/Weights/Linear.lean index bd48f9a203f78..f52cb53419a1b 100644 --- a/Mathlib/Algebra/Lie/Weights/Linear.lean +++ b/Mathlib/Algebra/Lie/Weights/Linear.lean @@ -48,14 +48,14 @@ namespace LieModule /-- A typeclass encoding the fact that a given Lie module has linear weights, vanishing on the derived ideal. -/ -class LinearWeights [LieAlgebra.IsNilpotent R L] : Prop where +class LinearWeights [LieRing.IsNilpotent L] : Prop where map_add : ∀ χ : L → R, genWeightSpace M χ ≠ ⊥ → ∀ x y, χ (x + y) = χ x + χ y map_smul : ∀ χ : L → R, genWeightSpace M χ ≠ ⊥ → ∀ (t : R) x, χ (t • x) = t • χ x map_lie : ∀ χ : L → R, genWeightSpace M χ ≠ ⊥ → ∀ x y : L, χ ⁅x, y⁆ = 0 namespace Weight -variable [LieAlgebra.IsNilpotent R L] [LinearWeights R L M] (χ : Weight R L M) +variable [LieRing.IsNilpotent L] [LinearWeights R L M] (χ : Weight R L M) /-- A weight of a Lie module, bundled as a linear map. -/ @[simps] @@ -115,7 +115,7 @@ section FiniteDimensional open Module variable [IsDomain R] [IsPrincipalIdealRing R] [Module.Free R M] [Module.Finite R M] - [LieAlgebra.IsNilpotent R L] + [LieRing.IsNilpotent L] lemma trace_comp_toEnd_genWeightSpace_eq (χ : L → R) : LinearMap.trace R _ ∘ₗ (toEnd R L (genWeightSpace M χ)).toLinearMap = @@ -154,7 +154,7 @@ instance instLinearWeightsOfCharZero [CharZero R] : end FiniteDimensional -variable [LieAlgebra.IsNilpotent R L] (χ : L → R) +variable [LieRing.IsNilpotent L] (χ : L → R) /-- A type synonym for the `χ`-weight space but with the action of `x : L` on `m : genWeightSpace M χ`, shifted to act as `⁅x, m⁆ - χ x • m`. -/ @@ -222,15 +222,15 @@ lemma toEnd_eq (x : L) : /-- By Engel's theorem, if `M` is Noetherian, the shifted action `⁅x, m⁆ - χ x • m` makes the `χ`-weight space into a nilpotent Lie module. -/ -instance [IsNoetherian R M] : IsNilpotent R L (shiftedGenWeightSpace R L M χ) := +instance [IsNoetherian R M] : IsNilpotent L (shiftedGenWeightSpace R L M χ) := LieModule.isNilpotent_iff_forall'.mpr fun x ↦ isNilpotent_toEnd_sub_algebraMap M χ x end shiftedGenWeightSpace open shiftedGenWeightSpace in -/-- Given a Lie module `M` of a Lie algebra `L` with coefficients in `R`, if a function `χ : L → R` -has a simultaneous generalized eigenvector for the action of `L` then it has a simultaneous true -eigenvector, provided `M` is Noetherian and has linear weights. -/ +/-- Given a Lie module `M` of a nilpotent Lie algebra `L` with coefficients in `R`, +if a function `χ : L → R` has a simultaneous generalized eigenvector for the action of `L` +then it has a simultaneous true eigenvector, provided `M` is Noetherian and has linear weights. -/ lemma exists_forall_lie_eq_smul [LinearWeights R L M] [IsNoetherian R M] (χ : Weight R L M) : ∃ m : M, m ≠ 0 ∧ ∀ x : L, ⁅x, m⁆ = χ x • m := by replace hχ : Nontrivial (shiftedGenWeightSpace R L M χ) := @@ -247,7 +247,7 @@ lemma exists_forall_lie_eq_smul [LinearWeights R L M] [IsNoetherian R M] (χ : W /-- See `LieModule.exists_nontrivial_weightSpace_of_isSolvable` for the variant that only assumes that `L` is solvable but additionally requires `k` to be of characteristic zero. -/ lemma exists_nontrivial_weightSpace_of_isNilpotent [Field k] [LieAlgebra k L] [Module k M] - [Module.Finite k M] [LieModule k L M] [LieAlgebra.IsNilpotent k L] [LinearWeights k L M] + [Module.Finite k M] [LieModule k L M] [LinearWeights k L M] [IsTriangularizable k L M] [Nontrivial M] : ∃ χ : Module.Dual k L, Nontrivial (weightSpace M χ) := by obtain ⟨χ⟩ : Nonempty (Weight k L M) := by diff --git a/Mathlib/Algebra/Module/Basic.lean b/Mathlib/Algebra/Module/Basic.lean index b653c02aa9f6c..7ba518660f634 100644 --- a/Mathlib/Algebra/Module/Basic.lean +++ b/Mathlib/Algebra/Module/Basic.lean @@ -3,12 +3,13 @@ Copyright (c) 2015 Nathaniel Thomas. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Nathaniel Thomas, Jeremy Avigad, Johannes Hölzl, Mario Carneiro -/ -import Mathlib.Algebra.Field.Basic +import Mathlib.Algebra.Field.Defs import Mathlib.Algebra.Group.Action.Pi import Mathlib.Algebra.Group.Indicator import Mathlib.Algebra.GroupWithZero.Action.Units import Mathlib.Algebra.Module.NatInt import Mathlib.Algebra.NoZeroSMulDivisors.Defs +import Mathlib.Algebra.Ring.Invertible /-! # Further basic results about modules. diff --git a/Mathlib/Algebra/Module/Equiv/Basic.lean b/Mathlib/Algebra/Module/Equiv/Basic.lean index a22205a1e0f06..b0228d7304c79 100644 --- a/Mathlib/Algebra/Module/Equiv/Basic.lean +++ b/Mathlib/Algebra/Module/Equiv/Basic.lean @@ -49,7 +49,7 @@ def restrictScalars (f : M ≃ₗ[S] M₂) : M ≃ₗ[R] M₂ := theorem restrictScalars_injective : Function.Injective (restrictScalars R : (M ≃ₗ[S] M₂) → M ≃ₗ[R] M₂) := fun _ _ h ↦ - ext (LinearEquiv.congr_fun h : _) + ext (LinearEquiv.congr_fun h :) @[simp] theorem restrictScalars_inj (f g : M ≃ₗ[S] M₂) : diff --git a/Mathlib/Algebra/Module/FinitePresentation.lean b/Mathlib/Algebra/Module/FinitePresentation.lean index 468b2cc116e18..f956657b81585 100644 --- a/Mathlib/Algebra/Module/FinitePresentation.lean +++ b/Mathlib/Algebra/Module/FinitePresentation.lean @@ -187,7 +187,7 @@ lemma Module.FinitePresentation.fg_ker [Module.Finite R M] exact ⟨_, hy, by simp⟩ apply Submodule.fg_of_fg_map_of_fg_inf_ker f.range.mkQ · rw [this] - exact Module.Finite.out + exact Module.Finite.fg_top · rw [Submodule.ker_mkQ, inf_comm, ← Submodule.map_comap_eq, ← LinearMap.ker_comp, hf] exact hs'.map f @@ -201,8 +201,8 @@ lemma Module.finitePresentation_of_ker [Module.FinitePresentation R N] Module.FinitePresentation R M := by obtain ⟨s, hs⟩ : (⊤ : Submodule R M).FG := by apply Submodule.fg_of_fg_map_of_fg_inf_ker l - · rw [Submodule.map_top, LinearMap.range_eq_top.mpr hl]; exact Module.Finite.out - · rw [top_inf_eq, ← Submodule.fg_top]; exact Module.Finite.out + · rw [Submodule.map_top, LinearMap.range_eq_top.mpr hl]; exact Module.Finite.fg_top + · rw [top_inf_eq, ← Submodule.fg_top]; exact Module.Finite.fg_top refine ⟨s, hs, ?_⟩ let π := Finsupp.linearCombination R ((↑) : s → M) have H : Function.Surjective π := @@ -250,7 +250,7 @@ instance {A} [CommRing A] [Algebra R A] [Module.FinitePresentation R M] : apply Submodule.FG.map have : Module.Finite R (LinearMap.ker f) := ⟨(Submodule.fg_top _).mpr (Module.FinitePresentation.fg_ker f hf)⟩ - exact Module.Finite.out (R := A) (M := A ⊗[R] LinearMap.ker f) + exact Module.Finite.fg_top (R := A) (M := A ⊗[R] LinearMap.ker f) open TensorProduct in lemma FinitePresentation.of_isBaseChange diff --git a/Mathlib/Algebra/Module/LinearMap/Defs.lean b/Mathlib/Algebra/Module/LinearMap/Defs.lean index d4c25f58c52fb..7041673c0691f 100644 --- a/Mathlib/Algebra/Module/LinearMap/Defs.lean +++ b/Mathlib/Algebra/Module/LinearMap/Defs.lean @@ -427,7 +427,7 @@ theorem restrictScalars_apply (fₗ : M →ₗ[S] M₂) (x) : restrictScalars R theorem restrictScalars_injective : Function.Injective (restrictScalars R : (M →ₗ[S] M₂) → M →ₗ[R] M₂) := fun _ _ h ↦ - ext (LinearMap.congr_fun h : _) + ext (LinearMap.congr_fun h :) @[simp] theorem restrictScalars_inj (fₗ gₗ : M →ₗ[S] M₂) : @@ -570,7 +570,7 @@ instance CompatibleSMul.intModule {S : Type*} [Semiring S] [Module S M] [Module instance CompatibleSMul.units {R S : Type*} [Monoid R] [MulAction R M] [MulAction R M₂] [Semiring S] [Module S M] [Module S M₂] [CompatibleSMul M M₂ R S] : CompatibleSMul M M₂ Rˣ S := - ⟨fun fₗ c x ↦ (CompatibleSMul.map_smul fₗ (c : R) x : _)⟩ + ⟨fun fₗ c x ↦ (CompatibleSMul.map_smul fₗ (c : R) x :)⟩ end AddCommGroup diff --git a/Mathlib/Algebra/Module/LinearMap/End.lean b/Mathlib/Algebra/Module/LinearMap/End.lean index 5f1a7fe6fd3d0..d7cac2488f1b0 100644 --- a/Mathlib/Algebra/Module/LinearMap/End.lean +++ b/Mathlib/Algebra/Module/LinearMap/End.lean @@ -89,7 +89,7 @@ theorem _root_.Module.End.natCast_apply (n : ℕ) (m : M) : (↑n : Module.End R @[simp] theorem _root_.Module.End.ofNat_apply (n : ℕ) [n.AtLeastTwo] (m : M) : - (no_index (OfNat.ofNat n) : Module.End R M) m = OfNat.ofNat n • m := rfl + (ofNat(n) : Module.End R M) m = OfNat.ofNat n • m := rfl instance _root_.Module.End.ring : Ring (Module.End R N₁) := { Module.End.semiring, LinearMap.addCommGroup with diff --git a/Mathlib/Algebra/Module/LocalizedModule/Basic.lean b/Mathlib/Algebra/Module/LocalizedModule/Basic.lean index 9603b7c878309..b44b61b98fe6a 100644 --- a/Mathlib/Algebra/Module/LocalizedModule/Basic.lean +++ b/Mathlib/Algebra/Module/LocalizedModule/Basic.lean @@ -677,8 +677,8 @@ there is a linear map `lift g ∘ mkLinearMap = g`. -/ theorem lift_comp (g : M →ₗ[R] M'') (h : ∀ x : S, IsUnit ((algebraMap R (Module.End R M'')) x)) : (lift S g h).comp (mkLinearMap S M) = g := by - ext x; dsimp; rw [LocalizedModule.lift_mk] - erw [Module.End_algebraMap_isUnit_inv_apply_eq_iff, one_smul] + ext x + simp [LocalizedModule.lift_mk] /-- If `g` is a linear map `M → M''` such that all scalar multiplication by `s : S` is invertible and @@ -1233,7 +1233,7 @@ theorem mkOfAlgebra {R S S' : Type*} [CommRing R] [CommRing S] [CommRing S'] [Al simp_rw [Submonoid.smul_def, Algebra.smul_def] at e exact (h₁ x x.2).mul_left_cancel e · intro a - refine ⟨((h₁ x x.2).unit⁻¹ : _) * a, ?_⟩ + refine ⟨((h₁ x x.2).unit⁻¹ :) * a, ?_⟩ rw [Module.algebraMap_end_apply, Algebra.smul_def, ← mul_assoc, IsUnit.mul_val_inv, one_mul] · exact h₂ · intros x y diff --git a/Mathlib/Algebra/Module/Pi.lean b/Mathlib/Algebra/Module/Pi.lean index 23c86123bb947..944e5906eb762 100644 --- a/Mathlib/Algebra/Module/Pi.lean +++ b/Mathlib/Algebra/Module/Pi.lean @@ -26,7 +26,7 @@ namespace Pi theorem _root_.IsSMulRegular.pi {α : Type*} [∀ i, SMul α <| f i] {k : α} (hk : ∀ i, IsSMulRegular (f i) k) : IsSMulRegular (∀ i, f i) k := fun _ _ h => - funext fun i => hk i (congr_fun h i : _) + funext fun i => hk i (congr_fun h i :) instance smulWithZero (α) [Zero α] [∀ i, Zero (f i)] [∀ i, SMulWithZero α (f i)] : SMulWithZero α (∀ i, f i) := diff --git a/Mathlib/Algebra/Module/Presentation/Basic.lean b/Mathlib/Algebra/Module/Presentation/Basic.lean index b12188e759ad7..e9c9735f6af10 100644 --- a/Mathlib/Algebra/Module/Presentation/Basic.lean +++ b/Mathlib/Algebra/Module/Presentation/Basic.lean @@ -354,8 +354,8 @@ certain linear equations. -/ noncomputable def linearMapEquiv : (M →ₗ[A] N) ≃ relations.Solution N where toFun f := solution.postcomp f invFun s := h.desc s - left_inv f := h.postcomp_injective (by aesop) - right_inv s := by aesop + left_inv f := h.postcomp_injective (by simp) + right_inv s := by simp section diff --git a/Mathlib/Algebra/Module/Submodule/Basic.lean b/Mathlib/Algebra/Module/Submodule/Basic.lean index e6ab0afbfd51c..1e3b2839af13e 100644 --- a/Mathlib/Algebra/Module/Submodule/Basic.lean +++ b/Mathlib/Algebra/Module/Submodule/Basic.lean @@ -99,7 +99,7 @@ instance [VAdd M α] : VAdd p α := p.toAddSubmonoid.vadd instance vaddCommClass [VAdd M β] [VAdd α β] [VAddCommClass M α β] : VAddCommClass p α β := - ⟨fun a => (vadd_comm (a : M) : _)⟩ + ⟨fun a => vadd_comm (a : M)⟩ instance [VAdd M α] [FaithfulVAdd M α] : FaithfulVAdd p α := ⟨fun h => Subtype.ext <| eq_of_vadd_eq_vadd h⟩ diff --git a/Mathlib/Algebra/Module/Submodule/Defs.lean b/Mathlib/Algebra/Module/Submodule/Defs.lean index 572cdc0af7b5b..985489a5a1fbe 100644 --- a/Mathlib/Algebra/Module/Submodule/Defs.lean +++ b/Mathlib/Algebra/Module/Submodule/Defs.lean @@ -288,7 +288,7 @@ theorem mem_toAddSubgroup : x ∈ p.toAddSubgroup ↔ x ∈ p := Iff.rfl theorem toAddSubgroup_injective : Injective (toAddSubgroup : Submodule R M → AddSubgroup M) - | _, _, h => SetLike.ext (SetLike.ext_iff.1 h : _) + | _, _, h => SetLike.ext (SetLike.ext_iff.1 h :) @[simp] theorem toAddSubgroup_inj : p.toAddSubgroup = p'.toAddSubgroup ↔ p = p' := diff --git a/Mathlib/Algebra/Module/Submodule/Pointwise.lean b/Mathlib/Algebra/Module/Submodule/Pointwise.lean index bad57cad693c0..811814d19e270 100644 --- a/Mathlib/Algebra/Module/Submodule/Pointwise.lean +++ b/Mathlib/Algebra/Module/Submodule/Pointwise.lean @@ -347,11 +347,6 @@ instance : CovariantClass (Set S) (Submodule R M) HSMul.hSMul LE.le := ⟨fun _ _ _ le => set_smul_le _ _ _ fun _ _ hr hm => mem_set_smul_of_mem_mem (mem1 := hr) (mem2 := le hm)⟩ -@[deprecated smul_mono_right (since := "2024-03-31")] -theorem set_smul_mono_right {p q : Submodule R M} (le : p ≤ q) : - s • p ≤ s • q := - smul_mono_right s le - lemma set_smul_mono_left {s t : Set S} (le : s ≤ t) : s • N ≤ t • N := set_smul_le _ _ _ fun _ _ hr hm => mem_set_smul_of_mem_mem (mem1 := le hr) @@ -475,8 +470,8 @@ lemma mem_singleton_set_smul [SMulCommClass R S M] (r : S) (x : M) : exact ⟨t • n, by aesop, smul_comm _ _ _⟩ · rcases h₁ with ⟨m₁, h₁, rfl⟩ rcases h₂ with ⟨m₂, h₂, rfl⟩ - exact ⟨m₁ + m₂, Submodule.add_mem _ h₁ h₂, by aesop⟩ - · exact ⟨0, Submodule.zero_mem _, by aesop⟩ + exact ⟨m₁ + m₂, Submodule.add_mem _ h₁ h₂, by simp⟩ + · exact ⟨0, Submodule.zero_mem _, by simp⟩ · aesop lemma smul_inductionOn_pointwise [SMulCommClass S R M] {a : S} {p : (x : M) → x ∈ a • N → Prop} diff --git a/Mathlib/Algebra/Module/Submodule/RestrictScalars.lean b/Mathlib/Algebra/Module/Submodule/RestrictScalars.lean index 54f609912e963..6cda15b8f5f30 100644 --- a/Mathlib/Algebra/Module/Submodule/RestrictScalars.lean +++ b/Mathlib/Algebra/Module/Submodule/RestrictScalars.lean @@ -55,7 +55,7 @@ variable (R M) theorem restrictScalars_injective : Function.Injective (restrictScalars S : Submodule R M → Submodule S M) := fun _ _ h => - ext <| Set.ext_iff.1 (SetLike.ext'_iff.1 h : _) + ext <| Set.ext_iff.1 (SetLike.ext'_iff.1 h :) @[simp] theorem restrictScalars_inj {V₁ V₂ : Submodule R M} : diff --git a/Mathlib/Algebra/Module/Torsion.lean b/Mathlib/Algebra/Module/Torsion.lean index 08b33bcef771a..7443d4cffa6a7 100644 --- a/Mathlib/Algebra/Module/Torsion.lean +++ b/Mathlib/Algebra/Module/Torsion.lean @@ -408,6 +408,8 @@ theorem supIndep_torsionBySet_ideal (hp : (S : Set ι).Pairwise fun i j => p i variable {q : ι → R} +open scoped Function -- required for scoped `on` notation + theorem iSup_torsionBy_eq_torsionBy_prod (hq : (S : Set ι).Pairwise <| (IsCoprime on q)) : ⨆ i ∈ S, torsionBy R M (q i) = torsionBy R M (∏ i ∈ S, q i) := by rw [← torsionBySet_span_singleton_eq, Ideal.submodule_span_eq, ← @@ -454,6 +456,7 @@ theorem torsionBySet_isInternal {p : ι → Ideal R} apply (iSup_torsionBySet_ideal_eq_torsionBySet_iInf hp).trans <| (Module.isTorsionBySet_iff_torsionBySet_eq_top _).mp hM) +open scoped Function in -- required for scoped `on` notation /-- If the `q i` are pairwise coprime, a `∏ i, q i`-torsion module is the internal direct sum of its `q i`-torsion submodules. -/ theorem torsionBy_isInternal {q : ι → R} (hq : (S : Set ι).Pairwise <| (IsCoprime on q)) @@ -501,7 +504,7 @@ instance IsTorsionBySet.isScalarTower (hM : IsTorsionBySet R M I) @IsScalarTower S (R ⧸ I) M _ (IsTorsionBySet.module hM).toSMul _ := -- Porting note: still needed to be fed the Module R / I M instance @IsScalarTower.mk S (R ⧸ I) M _ (IsTorsionBySet.module hM).toSMul _ - (fun b d x => Quotient.inductionOn' d fun c => (smul_assoc b c x : _)) + (fun b d x => Quotient.inductionOn' d fun c => (smul_assoc b c x :)) /-- An `(R ⧸ Ideal.span {r})`-module is an `R`-module for which `IsTorsionBy R M r`. -/ abbrev IsTorsionBy.module (hM : IsTorsionBy R M r) : Module (R ⧸ Ideal.span {r}) M := @@ -691,7 +694,7 @@ variable {R M} theorem _root_.Submodule.annihilator_top_inter_nonZeroDivisors [Module.Finite R M] (hM : Module.IsTorsion R M) : ((⊤ : Submodule R M).annihilator : Set R) ∩ R⁰ ≠ ∅ := by - obtain ⟨S, hS⟩ := ‹Module.Finite R M›.out + obtain ⟨S, hS⟩ := ‹Module.Finite R M›.fg_top refine Set.Nonempty.ne_empty ⟨_, ?_, (∏ x ∈ S, (@hM x).choose : R⁰).prop⟩ rw [Submonoid.coe_finset_prod, SetLike.mem_coe, ← hS, mem_annihilator_span] intro n diff --git a/Mathlib/Algebra/MonoidAlgebra/Basic.lean b/Mathlib/Algebra/MonoidAlgebra/Basic.lean index 1e0a8e01e33eb..a0db55e090c4b 100644 --- a/Mathlib/Algebra/MonoidAlgebra/Basic.lean +++ b/Mathlib/Algebra/MonoidAlgebra/Basic.lean @@ -369,7 +369,7 @@ def liftMagma [Module k A] [IsScalarTower k A A] [SMulCommClass k A A] : (Multiplicative G →ₙ* A) ≃ (k[G] →ₙₐ[k] A) := { (MonoidAlgebra.liftMagma k : (Multiplicative G →ₙ* A) ≃ (_ →ₙₐ[k] A)) with toFun := fun f => - { (MonoidAlgebra.liftMagma k f : _) with + { (MonoidAlgebra.liftMagma k f :) with toFun := fun a => sum a fun m t => t • f (Multiplicative.ofAdd m) } invFun := fun F => F.toMulHom.comp (ofMagma k G) } diff --git a/Mathlib/Algebra/MonoidAlgebra/Defs.lean b/Mathlib/Algebra/MonoidAlgebra/Defs.lean index da1c6d0c0a369..379ff57f2cd6b 100644 --- a/Mathlib/Algebra/MonoidAlgebra/Defs.lean +++ b/Mathlib/Algebra/MonoidAlgebra/Defs.lean @@ -993,7 +993,7 @@ theorem liftNC_mul {g_hom : Type*} (f : k →+* R) (g : g_hom) (a b : k[G]) (h_comm : ∀ {x y}, y ∈ a.support → Commute (f (b x)) (g <| Multiplicative.ofAdd y)) : liftNC (f : k →+ R) g (a * b) = liftNC (f : k →+ R) g a * liftNC (f : k →+ R) g b := - (MonoidAlgebra.liftNC_mul f g _ _ @h_comm : _) + MonoidAlgebra.liftNC_mul f g _ _ @h_comm end Mul @@ -1013,7 +1013,7 @@ theorem one_def : (1 : k[G]) = single 0 1 := theorem liftNC_one {g_hom : Type*} [FunLike g_hom (Multiplicative G) R] [OneHomClass g_hom (Multiplicative G) R] (f : k →+* R) (g : g_hom) : liftNC (f : k →+ R) g 1 = 1 := - (MonoidAlgebra.liftNC_one f g : _) + MonoidAlgebra.liftNC_one f g end One diff --git a/Mathlib/Algebra/MonoidAlgebra/ToDirectSum.lean b/Mathlib/Algebra/MonoidAlgebra/ToDirectSum.lean index 2d06dc8ac12c6..81771c479c9ed 100644 --- a/Mathlib/Algebra/MonoidAlgebra/ToDirectSum.lean +++ b/Mathlib/Algebra/MonoidAlgebra/ToDirectSum.lean @@ -96,7 +96,7 @@ theorem AddMonoidAlgebra.toDirectSum_toAddMonoidAlgebra (f : AddMonoidAlgebra M @[simp] theorem DirectSum.toAddMonoidAlgebra_toDirectSum (f : ⨁ _ : ι, M) : f.toAddMonoidAlgebra.toDirectSum = f := - (DFinsupp.toFinsupp_toDFinsupp (show Π₀ _ : ι, M from f) : _) + (DFinsupp.toFinsupp_toDFinsupp (show Π₀ _ : ι, M from f) :) end diff --git a/Mathlib/Algebra/MvPolynomial/Equiv.lean b/Mathlib/Algebra/MvPolynomial/Equiv.lean index 5080eada19bc1..444676431d13e 100644 --- a/Mathlib/Algebra/MvPolynomial/Equiv.lean +++ b/Mathlib/Algebra/MvPolynomial/Equiv.lean @@ -250,7 +250,7 @@ def sumAlgEquiv : MvPolynomial (S₁ ⊕ S₂) R ≃ₐ[R] MvPolynomial S₁ (Mv { sumRingEquiv R S₁ S₂ with commutes' := by intro r - have A : algebraMap R (MvPolynomial S₁ (MvPolynomial S₂ R)) r = (C (C r) : _) := rfl + have A : algebraMap R (MvPolynomial S₁ (MvPolynomial S₂ R)) r = (C (C r) :) := rfl have B : algebraMap R (MvPolynomial (S₁ ⊕ S₂) R) r = C r := rfl simp only [sumRingEquiv, mvPolynomialEquivMvPolynomial, Equiv.toFun_as_coe, Equiv.coe_fn_mk, B, sumToIter_C, A] } diff --git a/Mathlib/Algebra/Order/AbsoluteValue/Equivalence.lean b/Mathlib/Algebra/Order/AbsoluteValue/Equivalence.lean index bd4fa328898cc..f63cfd022b9c4 100644 --- a/Mathlib/Algebra/Order/AbsoluteValue/Equivalence.lean +++ b/Mathlib/Algebra/Order/AbsoluteValue/Equivalence.lean @@ -18,45 +18,45 @@ variable {R : Type*} [Semiring R] /-- Two absolute values `f, g` on `R` with values in `ℝ` are *equivalent* if there exists a positive real constant `c` such that for all `x : R`, `(f x)^c = g x`. -/ -def Equiv (f g : AbsoluteValue R ℝ) : Prop := +def IsEquiv (f g : AbsoluteValue R ℝ) : Prop := ∃ c : ℝ, 0 < c ∧ (f · ^ c) = g /-- Equivalence of absolute values is reflexive. -/ -lemma equiv_refl (f : AbsoluteValue R ℝ) : Equiv f f := +lemma isEquiv_refl (f : AbsoluteValue R ℝ) : IsEquiv f f := ⟨1, Real.zero_lt_one, funext fun x ↦ Real.rpow_one (f x)⟩ /-- Equivalence of absolute values is symmetric. -/ -lemma equiv_symm {f g : AbsoluteValue R ℝ} (hfg : Equiv f g) : Equiv g f := by +lemma isEquiv_symm {f g : AbsoluteValue R ℝ} (hfg : IsEquiv f g) : IsEquiv g f := by rcases hfg with ⟨c, hcpos, h⟩ refine ⟨1 / c, one_div_pos.mpr hcpos, ?_⟩ simp [← h, Real.rpow_rpow_inv (apply_nonneg f _) (ne_of_lt hcpos).symm] /-- Equivalence of absolute values is transitive. -/ -lemma equiv_trans {f g k : AbsoluteValue R ℝ} (hfg : Equiv f g) (hgk : Equiv g k) : - Equiv f k := by +lemma isEquiv_trans {f g k : AbsoluteValue R ℝ} (hfg : IsEquiv f g) (hgk : IsEquiv g k) : + IsEquiv f k := by rcases hfg with ⟨c, hcPos, hfg⟩ rcases hgk with ⟨d, hdPos, hgk⟩ refine ⟨c * d, mul_pos hcPos hdPos, ?_⟩ simp [← hgk, ← hfg, Real.rpow_mul (apply_nonneg f _)] +instance : Setoid (AbsoluteValue R ℝ) where + r := IsEquiv + iseqv := { + refl := isEquiv_refl + symm := isEquiv_symm + trans := isEquiv_trans + } + /-- An absolute value is equivalent to the trivial iff it is trivial itself. -/ @[simp] -lemma eq_trivial_of_equiv_trivial [DecidablePred fun x : R ↦ x = 0] [NoZeroDivisors R] +lemma eq_trivial_of_isEquiv_trivial [DecidablePred fun x : R ↦ x = 0] [NoZeroDivisors R] (f : AbsoluteValue R ℝ) : - f.Equiv .trivial ↔ f = .trivial := by - refine ⟨fun ⟨c, hc₀, hc⟩ ↦ ext fun x ↦ ?_, fun H ↦ H ▸ equiv_refl f⟩ + f ≈ .trivial ↔ f = .trivial := by + refine ⟨fun ⟨c, hc₀, hc⟩ ↦ ext fun x ↦ ?_, fun H ↦ H ▸ isEquiv_refl f⟩ apply_fun (· x) at hc rcases eq_or_ne x 0 with rfl | hx · simp · simp only [ne_eq, hx, not_false_eq_true, trivial_apply] at hc ⊢ exact (Real.rpow_left_inj (f.nonneg x) zero_le_one hc₀.ne').mp <| (Real.one_rpow c).symm ▸ hc -instance : Setoid (AbsoluteValue R ℝ) where - r := Equiv - iseqv := { - refl := equiv_refl - symm := equiv_symm - trans := equiv_trans - } - end AbsoluteValue diff --git a/Mathlib/Algebra/Order/BigOperators/Ring/Finset.lean b/Mathlib/Algebra/Order/BigOperators/Ring/Finset.lean index 607f6e70c1a6c..96a7184d1eef9 100644 --- a/Mathlib/Algebra/Order/BigOperators/Ring/Finset.lean +++ b/Mathlib/Algebra/Order/BigOperators/Ring/Finset.lean @@ -217,8 +217,6 @@ lemma IsAbsoluteValue.abv_sum [Semiring R] [OrderedSemiring S] (abv : R → S) [ (f : ι → R) (s : Finset ι) : abv (∑ i ∈ s, f i) ≤ ∑ i ∈ s, abv (f i) := (IsAbsoluteValue.toAbsoluteValue abv).sum_le _ _ -@[deprecated (since := "2024-02-14")] alias abv_sum_le_sum_abv := IsAbsoluteValue.abv_sum - nonrec lemma AbsoluteValue.map_prod [CommSemiring R] [Nontrivial R] [LinearOrderedCommRing S] (abv : AbsoluteValue R S) (f : ι → R) (s : Finset ι) : abv (∏ i ∈ s, f i) = ∏ i ∈ s, abv (f i) := diff --git a/Mathlib/Algebra/Order/Field/Basic.lean b/Mathlib/Algebra/Order/Field/Basic.lean index e65d0da3e42fb..d2b9446b43e01 100644 --- a/Mathlib/Algebra/Order/Field/Basic.lean +++ b/Mathlib/Algebra/Order/Field/Basic.lean @@ -166,15 +166,6 @@ theorem one_le_inv_iff : 1 ≤ a⁻¹ ↔ 0 < a ∧ a ≤ 1 := one_le_inv_iff₀ ### Relating two divisions. -/ -@[deprecated (since := "2024-02-16")] alias div_le_div_of_le_of_nonneg := div_le_div_of_nonneg_right -@[deprecated (since := "2024-02-16")] alias div_lt_div_of_lt := div_lt_div_of_pos_right -@[deprecated (since := "2024-02-16")] alias div_le_div_of_le_left := div_le_div_of_nonneg_left -@[deprecated (since := "2024-02-16")] alias div_lt_div_of_lt_left := div_lt_div_of_pos_left - -@[deprecated div_le_div_of_nonneg_right (since := "2024-02-16")] -lemma div_le_div_of_le (hc : 0 ≤ c) (hab : a ≤ b) : a / c ≤ b / c := - div_le_div_of_nonneg_right hab hc - @[deprecated div_le_div_iff_of_pos_right (since := "2024-11-12")] theorem div_le_div_right (hc : 0 < c) : a / c ≤ b / c ↔ a ≤ b := div_le_div_iff_of_pos_right hc diff --git a/Mathlib/Algebra/Order/Field/Power.lean b/Mathlib/Algebra/Order/Field/Power.lean index 0da30f4498bde..ac82917ae0852 100644 --- a/Mathlib/Algebra/Order/Field/Power.lean +++ b/Mathlib/Algebra/Order/Field/Power.lean @@ -55,8 +55,6 @@ theorem zpow_strictAnti (h₀ : 0 < a) (h₁ : a < 1) : StrictAnti (a ^ · : ℤ theorem zpow_lt_iff_lt (hx : 1 < a) : a ^ m < a ^ n ↔ m < n := zpow_lt_zpow_iff_right₀ hx -@[deprecated (since := "2024-02-10")] alias ⟨_, zpow_lt_of_lt⟩ := zpow_lt_iff_lt - @[deprecated zpow_le_zpow_iff_right₀ (since := "2024-10-08")] theorem zpow_le_iff_le (hx : 1 < a) : a ^ m ≤ a ^ n ↔ m ≤ n := zpow_le_zpow_iff_right₀ hx diff --git a/Mathlib/Algebra/Order/Group/Finset.lean b/Mathlib/Algebra/Order/Group/Finset.lean index e08078aa0b808..e3f518734c79d 100644 --- a/Mathlib/Algebra/Order/Group/Finset.lean +++ b/Mathlib/Algebra/Order/Group/Finset.lean @@ -14,8 +14,41 @@ import Mathlib.Data.Finset.Lattice.Fold # `Finset.sup` in a group -/ +open scoped Finset + assert_not_exists MonoidWithZero +namespace Multiset +variable {α : Type*} [DecidableEq α] + +@[simp] lemma toFinset_nsmul (s : Multiset α) : ∀ n ≠ 0, (n • s).toFinset = s.toFinset + | 0, h => by contradiction + | n + 1, _ => by + by_cases h : n = 0 + · rw [h, zero_add, one_nsmul] + · rw [add_nsmul, toFinset_add, one_nsmul, toFinset_nsmul s n h, Finset.union_idempotent] + +lemma toFinset_eq_singleton_iff (s : Multiset α) (a : α) : + s.toFinset = {a} ↔ card s ≠ 0 ∧ s = card s • {a} := by + refine ⟨fun H ↦ ⟨fun h ↦ ?_, ext' fun x ↦ ?_⟩, fun H ↦ ?_⟩ + · rw [card_eq_zero.1 h, toFinset_zero] at H + exact Finset.singleton_ne_empty _ H.symm + · rw [count_nsmul, count_singleton] + by_cases hx : x = a + · simp_rw [hx, ite_true, mul_one, count_eq_card] + intro y hy + rw [← mem_toFinset, H, Finset.mem_singleton] at hy + exact hy.symm + have hx' : x ∉ s := fun h' ↦ hx <| by rwa [← mem_toFinset, H, Finset.mem_singleton] at h' + simp_rw [count_eq_zero_of_not_mem hx', hx, ite_false, Nat.mul_zero] + simpa only [toFinset_nsmul _ _ H.1, toFinset_singleton] using congr($(H.2).toFinset) + +lemma toFinset_card_eq_one_iff (s : Multiset α) : + #s.toFinset = 1 ↔ Multiset.card s ≠ 0 ∧ ∃ a : α, s = Multiset.card s • {a} := by + simp_rw [Finset.card_eq_one, Multiset.toFinset_eq_singleton_iff, exists_and_left] + +end Multiset + namespace Finset variable {ι κ M G : Type*} diff --git a/Mathlib/Algebra/Order/Group/Multiset.lean b/Mathlib/Algebra/Order/Group/Multiset.lean new file mode 100644 index 0000000000000..970a962cfde9c --- /dev/null +++ b/Mathlib/Algebra/Order/Group/Multiset.lean @@ -0,0 +1,185 @@ +/- +Copyright (c) 2015 Microsoft Corporation. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Mario Carneiro +-/ +import Mathlib.Algebra.Group.Hom.Defs +import Mathlib.Algebra.Group.Nat.Basic +import Mathlib.Algebra.Order.Monoid.Unbundled.ExistsOfLE +import Mathlib.Algebra.Order.Sub.Defs +import Mathlib.Data.Multiset.Dedup + +/-! +# Multisets form an ordered monoid + +This file contains the ordered monoid instance on multisets, and lemmas related to it. + +See note [foundational algebra order theory]. +-/ + +open List Nat + +variable {α β : Type*} + +namespace Multiset + +/-! ### Additive monoid -/ + +instance instAddLeftMono : AddLeftMono (Multiset α) where elim _s _t _u := Multiset.add_le_add_left + +instance instAddLeftReflectLE : AddLeftReflectLE (Multiset α) where + elim _s _t _u := Multiset.le_of_add_le_add_left + +instance instAddCancelCommMonoid : AddCancelCommMonoid (Multiset α) where + add_comm := Multiset.add_comm + add_assoc := Multiset.add_assoc + zero_add := Multiset.zero_add + add_zero := Multiset.add_zero + add_left_cancel _ _ _ := Multiset.add_right_inj.1 + nsmul := nsmulRec + +lemma mem_of_mem_nsmul {a : α} {s : Multiset α} {n : ℕ} (h : a ∈ n • s) : a ∈ s := by + induction' n with n ih + · rw [zero_nsmul] at h + exact absurd h (not_mem_zero _) + · rw [succ_nsmul, mem_add] at h + exact h.elim ih id + +@[simp] +lemma mem_nsmul {a : α} {s : Multiset α} {n : ℕ} : a ∈ n • s ↔ n ≠ 0 ∧ a ∈ s := by + refine ⟨fun ha ↦ ⟨?_, mem_of_mem_nsmul ha⟩, fun h ↦ ?_⟩ + · rintro rfl + simp [zero_nsmul] at ha + obtain ⟨n, rfl⟩ := exists_eq_succ_of_ne_zero h.1 + rw [succ_nsmul, mem_add] + exact Or.inr h.2 + +lemma mem_nsmul_of_ne_zero {a : α} {s : Multiset α} {n : ℕ} (h0 : n ≠ 0) : a ∈ n • s ↔ a ∈ s := by + simp [*] + +lemma nsmul_cons {s : Multiset α} (n : ℕ) (a : α) : + n • (a ::ₘ s) = n • ({a} : Multiset α) + n • s := by + rw [← singleton_add, nsmul_add] + +/-! ### Cardinality -/ + +/-- `Multiset.card` bundled as a group hom. -/ +@[simps] +def cardHom : Multiset α →+ ℕ where + toFun := card + map_zero' := card_zero + map_add' := card_add + +@[simp] +lemma card_nsmul (s : Multiset α) (n : ℕ) : card (n • s) = n * card s := cardHom.map_nsmul .. + +/-! ### `Multiset.replicate` -/ + +/-- `Multiset.replicate` as an `AddMonoidHom`. -/ +@[simps] +def replicateAddMonoidHom (a : α) : ℕ →+ Multiset α where + toFun n := replicate n a + map_zero' := replicate_zero a + map_add' _ _ := replicate_add _ _ a + +lemma nsmul_replicate {a : α} (n m : ℕ) : n • replicate m a = replicate (n * m) a := + ((replicateAddMonoidHom a).map_nsmul _ _).symm + +lemma nsmul_singleton (a : α) (n) : n • ({a} : Multiset α) = replicate n a := by + rw [← replicate_one, nsmul_replicate, mul_one] + +/-! ### `Multiset.map` -/ + +/-- `Multiset.map` as an `AddMonoidHom`. -/ +@[simps] +def mapAddMonoidHom (f : α → β) : Multiset α →+ Multiset β where + toFun := map f + map_zero' := map_zero _ + map_add' := map_add _ + +@[simp] +lemma coe_mapAddMonoidHom (f : α → β) : (mapAddMonoidHom f : Multiset α → Multiset β) = map f := rfl + +lemma map_nsmul (f : α → β) (n : ℕ) (s) : map f (n • s) = n • map f s := + (mapAddMonoidHom f).map_nsmul _ _ + +/-! ### Subtraction -/ + +section +variable [DecidableEq α] + +instance : OrderedSub (Multiset α) where tsub_le_iff_right _n _m _k := Multiset.sub_le_iff_le_add + +instance : ExistsAddOfLE (Multiset α) where + exists_add_of_le h := leInductionOn h fun s ↦ + let ⟨l, p⟩ := s.exists_perm_append; ⟨l, Quot.sound p⟩ + +end + +/-! ### `Multiset.filter` -/ + +section +variable (p : α → Prop) [DecidablePred p] + +lemma filter_nsmul (s : Multiset α) (n : ℕ) : filter p (n • s) = n • filter p s := by + refine s.induction_on ?_ ?_ + · simp only [filter_zero, nsmul_zero] + · intro a ha ih + rw [nsmul_cons, filter_add, ih, filter_cons, nsmul_add] + congr + split_ifs with hp <;> + · simp only [filter_eq_self, nsmul_zero, filter_eq_nil] + intro b hb + rwa [mem_singleton.mp (mem_of_mem_nsmul hb)] + +/-! ### countP -/ + +@[simp] +lemma countP_nsmul (s) (n : ℕ) : countP p (n • s) = n * countP p s := by + induction n <;> simp [*, succ_nsmul, succ_mul, zero_nsmul] + +/-- `countP p`, the number of elements of a multiset satisfying `p`, promoted to an +`AddMonoidHom`. -/ +def countPAddMonoidHom : Multiset α →+ ℕ where + toFun := countP p + map_zero' := countP_zero _ + map_add' := countP_add _ + +@[simp] lemma coe_countPAddMonoidHom : (countPAddMonoidHom p : Multiset α → ℕ) = countP p := rfl + +end + +@[simp] lemma dedup_nsmul [DecidableEq α] {s : Multiset α} {n : ℕ} (hn : n ≠ 0) : + (n • s).dedup = s.dedup := by ext a; by_cases h : a ∈ s <;> simp [h, hn] + +lemma Nodup.le_nsmul_iff_le {s t : Multiset α} {n : ℕ} (h : s.Nodup) (hn : n ≠ 0) : + s ≤ n • t ↔ s ≤ t := by + classical simp [← h.le_dedup_iff_le, Iff.comm, ← h.le_dedup_iff_le, hn] + +/-! ### Multiplicity of an element -/ + +section +variable [DecidableEq α] {s : Multiset α} + +/-- `count a`, the multiplicity of `a` in a multiset, promoted to an `AddMonoidHom`. -/ +def countAddMonoidHom (a : α) : Multiset α →+ ℕ := + countPAddMonoidHom (a = ·) + +@[simp] +lemma coe_countAddMonoidHom (a : α) : (countAddMonoidHom a : Multiset α → ℕ) = count a := rfl + +@[simp] +lemma count_nsmul (a : α) (n s) : count a (n • s) = n * count a s := by + induction n <;> simp [*, succ_nsmul, succ_mul, zero_nsmul] + +end + +-- TODO: This should be `addMonoidHom_ext` +@[ext] +lemma addHom_ext [AddZeroClass β] ⦃f g : Multiset α →+ β⦄ (h : ∀ x, f {x} = g {x}) : f = g := by + ext s + induction' s using Multiset.induction_on with a s ih + · simp only [_root_.map_zero] + · simp only [← singleton_add, _root_.map_add, ih, h] + +end Multiset diff --git a/Mathlib/Algebra/Order/Group/Pointwise/Interval.lean b/Mathlib/Algebra/Order/Group/Pointwise/Interval.lean index fb5b4cae02b8a..8a43aa8f926fc 100644 --- a/Mathlib/Algebra/Order/Group/Pointwise/Interval.lean +++ b/Mathlib/Algebra/Order/Group/Pointwise/Interval.lean @@ -778,15 +778,22 @@ theorem image_mul_left_Ioc {a : α} (h : 0 < a) (b c : α) : /-- The (pre)image under `inv` of `Ioo 0 a` is `Ioi a⁻¹`. -/ theorem inv_Ioo_0_left {a : α} (ha : 0 < a) : (Ioo 0 a)⁻¹ = Ioi a⁻¹ := by ext x - exact - ⟨fun h => inv_inv x ▸ (inv_lt_inv₀ ha h.1).2 h.2, fun h => - ⟨inv_pos.2 <| (inv_pos.2 ha).trans h, - inv_inv a ▸ (inv_lt_inv₀ ((inv_pos.2 ha).trans h) - (inv_pos.2 ha)).2 h⟩⟩ + exact ⟨fun h ↦ inv_lt_of_inv_lt₀ (inv_pos.1 h.1) h.2, + fun h ↦ ⟨inv_pos.2 <| (inv_pos.2 ha).trans h, inv_lt_of_inv_lt₀ ha h⟩⟩ + +/-- The (pre)image under `inv` of `Ioo a 0` is `Iio a⁻¹`. -/ +theorem inv_Ioo_0_right {a : α} (ha : a < 0) : (Ioo a 0)⁻¹ = Iio a⁻¹ := by + ext x + refine ⟨fun h ↦ (lt_inv_of_neg (inv_neg''.1 h.2) ha).2 h.1, fun h ↦ ?_⟩ + have h' := (h.trans (inv_neg''.2 ha)) + exact ⟨(lt_inv_of_neg ha h').2 h, inv_neg''.2 h'⟩ theorem inv_Ioi₀ {a : α} (ha : 0 < a) : (Ioi a)⁻¹ = Ioo 0 a⁻¹ := by rw [inv_eq_iff_eq_inv, inv_Ioo_0_left (inv_pos.2 ha), inv_inv] +theorem inv_Iio₀ {a : α} (ha : a < 0) : (Iio a)⁻¹ = Ioo a⁻¹ 0 := by + rw [inv_eq_iff_eq_inv, inv_Ioo_0_right (inv_neg''.2 ha), inv_inv] + theorem image_const_mul_Ioi_zero {k : Type*} [LinearOrderedField k] {x : k} (hx : 0 < x) : (fun y => x * y) '' Ioi (0 : k) = Ioi 0 := by erw [(Units.mk0 x hx.ne').mulLeft.image_eq_preimage, diff --git a/Mathlib/Algebra/Order/Interval/Set/Group.lean b/Mathlib/Algebra/Order/Interval/Set/Group.lean index 671a053591075..185f1f2c53353 100644 --- a/Mathlib/Algebra/Order/Interval/Set/Group.lean +++ b/Mathlib/Algebra/Order/Interval/Set/Group.lean @@ -146,6 +146,7 @@ end LinearOrderedAddCommGroup /-! ### Lemmas about disjointness of translates of intervals -/ +open scoped Function -- required for scoped `on` notation section PairwiseDisjoint section OrderedCommGroup diff --git a/Mathlib/Algebra/Order/Interval/Set/Instances.lean b/Mathlib/Algebra/Order/Interval/Set/Instances.lean index 4624971afef4b..06d73380d901f 100644 --- a/Mathlib/Algebra/Order/Interval/Set/Instances.lean +++ b/Mathlib/Algebra/Order/Interval/Set/Instances.lean @@ -267,9 +267,9 @@ instance cancelMonoid {α : Type*} [StrictOrderedRing α] [IsDomain α] : CancelMonoid (Ioc (0 : α) 1) := { Set.Ioc.monoid with mul_left_cancel := fun a _ _ h => - Subtype.ext <| mul_left_cancel₀ a.prop.1.ne' <| (congr_arg Subtype.val h : _) + Subtype.ext <| mul_left_cancel₀ a.prop.1.ne' <| (congr_arg Subtype.val h :) mul_right_cancel := fun _ b _ h => - Subtype.ext <| mul_right_cancel₀ b.prop.1.ne' <| (congr_arg Subtype.val h : _) } + Subtype.ext <| mul_right_cancel₀ b.prop.1.ne' <| (congr_arg Subtype.val h :) } instance cancelCommMonoid {α : Type*} [StrictOrderedCommRing α] [IsDomain α] : CancelCommMonoid (Ioc (0 : α) 1) := diff --git a/Mathlib/Algebra/Order/Kleene.lean b/Mathlib/Algebra/Order/Kleene.lean index 208f87b231cde..96c7b6b8be108 100644 --- a/Mathlib/Algebra/Order/Kleene.lean +++ b/Mathlib/Algebra/Order/Kleene.lean @@ -133,6 +133,14 @@ scoped[Computability] attribute [simp] add_eq_sup theorem add_idem (a : α) : a + a = a := by simp +lemma natCast_eq_one {n : ℕ} (nezero : n ≠ 0) : (n : α) = 1 := by + induction n, Nat.one_le_iff_ne_zero.mpr nezero using Nat.le_induction with + | base => exact Nat.cast_one + | succ x _ hx => rw [Nat.cast_add, hx, Nat.cast_one, add_idem 1] + +lemma ofNat_eq_one {n : ℕ} [n.AtLeastTwo] : (ofNat(n) : α) = 1 := + natCast_eq_one <| Nat.not_eq_zero_of_lt Nat.AtLeastTwo.prop + theorem nsmul_eq_self : ∀ {n : ℕ} (_ : n ≠ 0) (a : α), n • a = a | 0, h => (h rfl).elim | 1, _ => one_nsmul diff --git a/Mathlib/Algebra/Order/Ring/Unbundled/Rat.lean b/Mathlib/Algebra/Order/Ring/Unbundled/Rat.lean index 21b69f6811d62..bd7669e3210be 100644 --- a/Mathlib/Algebra/Order/Ring/Unbundled/Rat.lean +++ b/Mathlib/Algebra/Order/Ring/Unbundled/Rat.lean @@ -197,9 +197,6 @@ instance : AddLeftMono ℚ where @[simp] lemma num_pos {a : ℚ} : 0 < a.num ↔ 0 < a := lt_iff_lt_of_le_iff_le num_nonpos @[simp] lemma num_neg {a : ℚ} : a.num < 0 ↔ a < 0 := lt_iff_lt_of_le_iff_le num_nonneg -@[deprecated (since := "2024-02-16")] alias num_nonneg_iff_zero_le := num_nonneg -@[deprecated (since := "2024-02-16")] alias num_pos_iff_pos := num_pos - theorem div_lt_div_iff_mul_lt_mul {a b c d : ℤ} (b_pos : 0 < b) (d_pos : 0 < d) : (a : ℚ) / b < c / d ↔ a * d < c * b := by simp only [lt_iff_le_not_le] diff --git a/Mathlib/Algebra/Order/SuccPred/WithBot.lean b/Mathlib/Algebra/Order/SuccPred/WithBot.lean index df9812d972bf8..08d2abbd2dca9 100644 --- a/Mathlib/Algebra/Order/SuccPred/WithBot.lean +++ b/Mathlib/Algebra/Order/SuccPred/WithBot.lean @@ -24,6 +24,6 @@ lemma succ_one : succ (1 : WithBot α) = 2 := by simpa [one_add_one_eq_two] usin @[simp] lemma succ_ofNat (n : ℕ) [n.AtLeastTwo] : - succ (no_index (OfNat.ofNat n) : WithBot α) = OfNat.ofNat n + 1 := succ_natCast n + succ (ofNat(n) : WithBot α) = OfNat.ofNat n + 1 := succ_natCast n end WithBot diff --git a/Mathlib/Algebra/Polynomial/Degree/Lemmas.lean b/Mathlib/Algebra/Polynomial/Degree/Lemmas.lean index c6557963d2e95..e007ecda8074d 100644 --- a/Mathlib/Algebra/Polynomial/Degree/Lemmas.lean +++ b/Mathlib/Algebra/Polynomial/Degree/Lemmas.lean @@ -180,6 +180,8 @@ theorem coeff_add_eq_right_of_lt (pn : p.natDegree < n) : (p + q).coeff n = q.co rw [add_comm] exact coeff_add_eq_left_of_lt pn +open scoped Function -- required for scoped `on` notation + theorem degree_sum_eq_of_disjoint (f : S → R[X]) (s : Finset S) (h : Set.Pairwise { i | i ∈ s ∧ f i ≠ 0 } (Ne on degree ∘ f)) : degree (s.sum f) = s.sup fun i => degree (f i) := by diff --git a/Mathlib/Algebra/Polynomial/Derivative.lean b/Mathlib/Algebra/Polynomial/Derivative.lean index c23fdfc0c1e98..e0154c45d7370 100644 --- a/Mathlib/Algebra/Polynomial/Derivative.lean +++ b/Mathlib/Algebra/Polynomial/Derivative.lean @@ -195,7 +195,7 @@ alias derivative_nat_cast := derivative_natCast @[simp] theorem derivative_ofNat (n : ℕ) [n.AtLeastTwo] : - derivative (no_index (OfNat.ofNat n) : R[X]) = 0 := + derivative (ofNat(n) : R[X]) = 0 := derivative_natCast theorem iterate_derivative_eq_zero {p : R[X]} {x : ℕ} (hx : p.natDegree < x) : diff --git a/Mathlib/Algebra/Polynomial/Div.lean b/Mathlib/Algebra/Polynomial/Div.lean index 3bcbfbaac6a50..48d16d1d4a6f2 100644 --- a/Mathlib/Algebra/Polynomial/Div.lean +++ b/Mathlib/Algebra/Polynomial/Div.lean @@ -396,7 +396,6 @@ theorem modByMonic_eq_zero_iff_dvd (hq : Monic q) : p %ₘ q = 0 ↔ q ∣ p := degree_eq_natDegree (mt leadingCoeff_eq_zero.2 hrpq0)] at this exact not_lt_of_ge (Nat.le_add_right _ _) (WithBot.coe_lt_coe.1 this)⟩ -@[deprecated (since := "2024-03-23")] alias dvd_iff_modByMonic_eq_zero := modByMonic_eq_zero_iff_dvd /-- See `Polynomial.mul_left_modByMonic` for the other multiplication order. That version, unlike this one, requires commutativity. -/ diff --git a/Mathlib/Algebra/Polynomial/Eval/Defs.lean b/Mathlib/Algebra/Polynomial/Eval/Defs.lean index af51045eb7816..5b47a0237e819 100644 --- a/Mathlib/Algebra/Polynomial/Eval/Defs.lean +++ b/Mathlib/Algebra/Polynomial/Eval/Defs.lean @@ -102,10 +102,9 @@ theorem eval₂_natCast (n : ℕ) : (n : R[X]).eval₂ f x = n := by @[deprecated (since := "2024-04-17")] alias eval₂_nat_cast := eval₂_natCast --- See note [no_index around OfNat.ofNat] @[simp] lemma eval₂_ofNat {S : Type*} [Semiring S] (n : ℕ) [n.AtLeastTwo] (f : R →+* S) (a : S) : - (no_index (OfNat.ofNat n : R[X])).eval₂ f a = OfNat.ofNat n := by + (ofNat(n) : R[X]).eval₂ f a = OfNat.ofNat n := by simp [OfNat.ofNat] variable [Semiring T] @@ -263,10 +262,9 @@ theorem eval₂_at_natCast {S : Type*} [Semiring S] (f : R →+* S) (n : ℕ) : @[deprecated (since := "2024-04-17")] alias eval₂_at_nat_cast := eval₂_at_natCast --- See note [no_index around OfNat.ofNat] @[simp] theorem eval₂_at_ofNat {S : Type*} [Semiring S] (f : R →+* S) (n : ℕ) [n.AtLeastTwo] : - p.eval₂ f (no_index (OfNat.ofNat n)) = f (p.eval (OfNat.ofNat n)) := by + p.eval₂ f ofNat(n) = f (p.eval (OfNat.ofNat n)) := by simp [OfNat.ofNat] @[simp] @@ -279,10 +277,9 @@ theorem eval_natCast {n : ℕ} : (n : R[X]).eval x = n := by simp only [← C_eq @[deprecated (since := "2024-04-17")] alias eval_nat_cast := eval_natCast --- See note [no_index around OfNat.ofNat] @[simp] lemma eval_ofNat (n : ℕ) [n.AtLeastTwo] (a : R) : - (no_index (OfNat.ofNat n : R[X])).eval a = OfNat.ofNat n := by + (ofNat(n) : R[X]).eval a = OfNat.ofNat n := by simp only [OfNat.ofNat, eval_natCast] @[simp] @@ -398,7 +395,7 @@ theorem natCast_comp {n : ℕ} : (n : R[X]).comp p = n := by rw [← C_eq_natCas alias nat_cast_comp := natCast_comp @[simp] -theorem ofNat_comp (n : ℕ) [n.AtLeastTwo] : (no_index (OfNat.ofNat n) : R[X]).comp p = n := +theorem ofNat_comp (n : ℕ) [n.AtLeastTwo] : (ofNat(n) : R[X]).comp p = n := natCast_comp @[simp] @@ -550,10 +547,9 @@ protected theorem map_natCast (n : ℕ) : (n : R[X]).map f = n := @[deprecated (since := "2024-04-17")] alias map_nat_cast := map_natCast --- See note [no_index around OfNat.ofNat] @[simp] protected theorem map_ofNat (n : ℕ) [n.AtLeastTwo] : - (no_index (OfNat.ofNat n) : R[X]).map f = OfNat.ofNat n := + (ofNat(n) : R[X]).map f = OfNat.ofNat n := show (n : R[X]).map f = n by rw [Polynomial.map_natCast] --TODO rename to `map_dvd_map` diff --git a/Mathlib/Algebra/Polynomial/FieldDivision.lean b/Mathlib/Algebra/Polynomial/FieldDivision.lean index 36d71e8b2b9ee..07d250cccea6d 100644 --- a/Mathlib/Algebra/Polynomial/FieldDivision.lean +++ b/Mathlib/Algebra/Polynomial/FieldDivision.lean @@ -3,6 +3,7 @@ Copyright (c) 2018 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Johannes Hölzl, Kim Morrison, Jens Wagemaker -/ +import Mathlib.Algebra.Order.Group.Finset import Mathlib.Algebra.Polynomial.Derivative import Mathlib.Algebra.Polynomial.Eval.SMul import Mathlib.Algebra.Polynomial.Roots diff --git a/Mathlib/Algebra/Polynomial/RingDivision.lean b/Mathlib/Algebra/Polynomial/RingDivision.lean index cf24ecbffc9f9..eb2061427a125 100644 --- a/Mathlib/Algebra/Polynomial/RingDivision.lean +++ b/Mathlib/Algebra/Polynomial/RingDivision.lean @@ -219,6 +219,7 @@ theorem isCoprime_X_sub_C_of_isUnit_sub {R} [CommRing R] {a b : R} (h : IsUnit ( congr exact h.val_inv_mul⟩ +open scoped Function in -- required for scoped `on` notation theorem pairwise_coprime_X_sub_C {K} [Field K] {I : Type v} {s : I → K} (H : Function.Injective s) : Pairwise (IsCoprime on fun i : I => X - C (s i)) := fun _ _ hij => isCoprime_X_sub_C_of_isUnit_sub (sub_ne_zero_of_ne <| H.ne hij).isUnit diff --git a/Mathlib/Algebra/Ring/Basic.lean b/Mathlib/Algebra/Ring/Basic.lean index 1682f8c413b9e..61ec82eae38d6 100644 --- a/Mathlib/Algebra/Ring/Basic.lean +++ b/Mathlib/Algebra/Ring/Basic.lean @@ -183,3 +183,43 @@ lemma noZeroDivisors_iff_isDomain_or_subsingleton [Ring α] : rw [← isCancelMulZero_iff_noZeroDivisors, isCancelMulZero_iff_isDomain_or_subsingleton] end NoZeroDivisors + +section DivisionMonoid +variable [DivisionMonoid R] [HasDistribNeg R] {a b : R} + +lemma one_div_neg_one_eq_neg_one : (1 : R) / -1 = -1 := + have : -1 * -1 = (1 : R) := by rw [neg_mul_neg, one_mul] + Eq.symm (eq_one_div_of_mul_eq_one_right this) + +lemma one_div_neg_eq_neg_one_div (a : R) : 1 / -a = -(1 / a) := + calc + 1 / -a = 1 / (-1 * a) := by rw [neg_eq_neg_one_mul] + _ = 1 / a * (1 / -1) := by rw [one_div_mul_one_div_rev] + _ = 1 / a * -1 := by rw [one_div_neg_one_eq_neg_one] + _ = -(1 / a) := by rw [mul_neg, mul_one] + +lemma div_neg_eq_neg_div (a b : R) : b / -a = -(b / a) := + calc + b / -a = b * (1 / -a) := by rw [← inv_eq_one_div, division_def] + _ = b * -(1 / a) := by rw [one_div_neg_eq_neg_one_div] + _ = -(b * (1 / a)) := by rw [neg_mul_eq_mul_neg] + _ = -(b / a) := by rw [mul_one_div] + +lemma neg_div (a b : R) : -b / a = -(b / a) := by + rw [neg_eq_neg_one_mul, mul_div_assoc, ← neg_eq_neg_one_mul] + +@[field_simps] +lemma neg_div' (a b : R) : -(b / a) = -b / a := by simp [neg_div] + +@[simp] +lemma neg_div_neg_eq (a b : R) : -a / -b = a / b := by rw [div_neg_eq_neg_div, neg_div, neg_neg] + +lemma neg_inv : -a⁻¹ = (-a)⁻¹ := by rw [inv_eq_one_div, inv_eq_one_div, div_neg_eq_neg_div] + +lemma div_neg (a : R) : a / -b = -(a / b) := by rw [← div_neg_eq_neg_div] + +lemma inv_neg : (-a)⁻¹ = -a⁻¹ := by rw [neg_inv] + +lemma inv_neg_one : (-1 : R)⁻¹ = -1 := by rw [← neg_inv, inv_one] + +end DivisionMonoid diff --git a/Mathlib/Algebra/Ring/Equiv.lean b/Mathlib/Algebra/Ring/Equiv.lean index f180f913b77df..f65b840107c50 100644 --- a/Mathlib/Algebra/Ring/Equiv.lean +++ b/Mathlib/Algebra/Ring/Equiv.lean @@ -530,6 +530,14 @@ theorem coe_prodCongr {R R' S S' : Type*} [NonUnitalNonAssocSemiring R] ⇑(RingEquiv.prodCongr f g) = Prod.map f g := rfl +/-- This is `Equiv.piOptionEquivProd` as a `RingEquiv`. -/ +@[simps!] +def piOptionEquivProd {ι : Type*} {R : Option ι → Type*} [Π i, NonUnitalNonAssocSemiring (R i)] : + (Π i, R i) ≃+* R none × (Π i, R (some i)) where + toEquiv := Equiv.piOptionEquivProd + map_add' _ _ := rfl + map_mul' _ _ := rfl + end NonUnitalSemiring section Semiring diff --git a/Mathlib/Algebra/Ring/Semireal/Defs.lean b/Mathlib/Algebra/Ring/Semireal/Defs.lean index 7b16914e15ab1..cb5317abf618d 100644 --- a/Mathlib/Algebra/Ring/Semireal/Defs.lean +++ b/Mathlib/Algebra/Ring/Semireal/Defs.lean @@ -8,12 +8,11 @@ import Mathlib.Algebra.Ring.SumsOfSquares /-! # Semireal rings -A semireal ring is a non-trivial commutative ring (with unit) in which `-1` is *not* a sum of -squares. Note that `-1` does not make sense in a semiring. +A semireal ring is a non-trivial commutative ring (with unit) in which `-1` is *not* a +sum of squares. -For instance, linearly ordered fields are semireal, because sums of squares are positive and `-1` is -not. More generally, linearly ordered semirings in which `a ≤ b → ∃ c, a + c = b` holds, are -semireal. +For instance, linearly ordered rings are semireal, because sums of squares are positive and `-1` is +not. ## Main declaration @@ -29,8 +28,8 @@ variable (R : Type*) /-- A semireal ring is a non-trivial commutative ring (with unit) in which `-1` is *not* a sum of -squares. Note that `-1` does not make sense in a semiring. Below we define the class `IsSemireal R` -for all additive monoid `R` equipped with a multiplication, a multiplicative unit and a negation. +squares. We define the class `IsSemireal R` +for all additive monoids `R` equipped with a multiplication, a multiplicative unit and a negation. -/ @[mk_iff] class IsSemireal [AddMonoid R] [Mul R] [One R] [Neg R] : Prop where diff --git a/Mathlib/Algebra/Ring/Subring/Defs.lean b/Mathlib/Algebra/Ring/Subring/Defs.lean index 3b201ab0ecd61..321e09cc9aa05 100644 --- a/Mathlib/Algebra/Ring/Subring/Defs.lean +++ b/Mathlib/Algebra/Ring/Subring/Defs.lean @@ -199,13 +199,13 @@ theorem copy_eq (S : Subring R) (s : Set R) (hs : s = ↑S) : S.copy s hs = S := SetLike.coe_injective hs theorem toSubsemiring_injective : Function.Injective (toSubsemiring : Subring R → Subsemiring R) - | _, _, h => ext (SetLike.ext_iff.mp h : _) + | _, _, h => ext (SetLike.ext_iff.mp h :) theorem toAddSubgroup_injective : Function.Injective (toAddSubgroup : Subring R → AddSubgroup R) - | _, _, h => ext (SetLike.ext_iff.mp h : _) + | _, _, h => ext (SetLike.ext_iff.mp h :) theorem toSubmonoid_injective : Function.Injective (fun s : Subring R => s.toSubmonoid) - | _, _, h => ext (SetLike.ext_iff.mp h : _) + | _, _, h => ext (SetLike.ext_iff.mp h :) /-- Construct a `Subring R` from a set `s`, a submonoid `sm`, and an additive subgroup `sa` such that `x ∈ s ↔ x ∈ sm ↔ x ∈ sa`. -/ diff --git a/Mathlib/Algebra/Ring/Subsemiring/Defs.lean b/Mathlib/Algebra/Ring/Subsemiring/Defs.lean index 3fe12b9558589..ada342ff44bd3 100644 --- a/Mathlib/Algebra/Ring/Subsemiring/Defs.lean +++ b/Mathlib/Algebra/Ring/Subsemiring/Defs.lean @@ -32,7 +32,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 + ofNat(n) ∈ s := by rw [← Nat.cast_ofNat]; exact natCast_mem s n instance (priority := 74) AddSubmonoidWithOneClass.toAddMonoidWithOne @@ -161,11 +161,11 @@ theorem copy_eq (S : Subsemiring R) (s : Set R) (hs : s = ↑S) : S.copy s hs = SetLike.coe_injective hs theorem toSubmonoid_injective : Function.Injective (toSubmonoid : Subsemiring R → Submonoid R) - | _, _, h => ext (SetLike.ext_iff.mp h : _) + | _, _, h => ext (SetLike.ext_iff.mp h :) theorem toAddSubmonoid_injective : Function.Injective (toAddSubmonoid : Subsemiring R → AddSubmonoid R) - | _, _, h => ext (SetLike.ext_iff.mp h : _) + | _, _, h => ext (SetLike.ext_iff.mp h :) /-- Construct a `Subsemiring R` from a set `s`, a submonoid `sm`, and an additive submonoid `sa` such that `x ∈ s ↔ x ∈ sm ↔ x ∈ sa`. -/ diff --git a/Mathlib/Algebra/Ring/SumsOfSquares.lean b/Mathlib/Algebra/Ring/SumsOfSquares.lean index 83f2856af82f4..cd768870f9db6 100644 --- a/Mathlib/Algebra/Ring/SumsOfSquares.lean +++ b/Mathlib/Algebra/Ring/SumsOfSquares.lean @@ -10,112 +10,111 @@ import Mathlib.Algebra.Order.Ring.Defs /-! # Sums of squares -We introduce sums of squares in a type `R` endowed with an `[Add R]`, `[Zero R]` and `[Mul R]` -instances. Sums of squares in `R` are defined by an inductive predicate `IsSumSq : R → Prop`: -`0 : R` is a sum of squares and if `S` is a sum of squares, then for all `a : R`, `a * a + S` is a -sum of squares in `R`. +We introduce a predicate for sums of squares in a ring. -## Main declaration +## Main declarations -- The predicate `IsSumSq : R → Prop`, defining the property of being a sum of squares in `R`. - -## Auxiliary declarations - -- The type `sumSqIn R` : in an additive monoid with multiplication `R`, we introduce the type -`sumSqIn R` as the submonoid of `R` whose carrier is the subset `{S : R | IsSumSq S}`. - -## Theorems - -- `IsSumSq.add`: if `S1` and `S2` are sums of squares in `R`, then so is `S1 + S2`. -- `IsSumSq.nonneg`: in a linearly ordered semiring `R` with an `[ExistsAddOfLE R]` instance, sums of -squares are non-negative. -- `mem_sumSqIn_of_isSquare`: a square is a sum of squares. -- `AddSubmonoid.closure_isSquare`: the submonoid of `R` generated by the squares in `R` is the set -of sums of squares in `R`. +- `IsSumSq : R → Prop`: for a type `R` with addition, multiplication and a zero, + an inductive predicate defining the property of being a sum of squares in `R`. + `0 : R` is a sum of squares and if `S` is a sum of squares, then, for all `a : R`, + `a * a + s` is a sum of squares. +- `AddMonoid.sumSq R`: the submonoid of sums of squares in an additive monoid `R` + with multiplication. -/ -variable {R : Type*} [Mul R] +variable {R : Type*} /-- -In a type `R` with an addition, a zero element and a multiplication, the property of being a sum of -squares is defined by an inductive predicate: `0 : R` is a sum of squares and if `S` is a sum of -squares, then for all `a : R`, `a * a + S` is a sum of squares in `R`. +The property of being a sum of squares is defined inductively by: +`0 : R` is a sum of squares and if `s : R` is a sum of squares, +then for all `a : R`, `a * a + s` is a sum of squares in `R`. -/ @[mk_iff] -inductive IsSumSq [Add R] [Zero R] : R → Prop - | zero : IsSumSq 0 - | sq_add (a S : R) (pS : IsSumSq S) : IsSumSq (a * a + S) +inductive IsSumSq [Mul R] [Add R] [Zero R] : R → Prop + | zero : IsSumSq 0 + | sq_add (a : R) {s : R} (hs : IsSumSq s) : IsSumSq (a * a + s) @[deprecated (since := "2024-08-09")] alias isSumSq := IsSumSq /-- -If `S1` and `S2` are sums of squares in a semiring `R`, then `S1 + S2` is a sum of squares in `R`. +In an additive monoid with multiplication, +if `s₁` and `s₂` are sums of squares, then `s₁ + s₂` is a sum of squares. -/ -theorem IsSumSq.add [AddMonoid R] {S1 S2 : R} (p1 : IsSumSq S1) - (p2 : IsSumSq S2) : IsSumSq (S1 + S2) := by - induction p1 with - | zero => rw [zero_add]; exact p2 - | sq_add a S pS ih => rw [add_assoc]; exact IsSumSq.sq_add a (S + S2) ih +theorem IsSumSq.add [AddMonoid R] [Mul R] {s₁ s₂ : R} + (h₁ : IsSumSq s₁) (h₂ : IsSumSq s₂) : IsSumSq (s₁ + s₂) := by + induction h₁ <;> simp_all [add_assoc, sq_add] @[deprecated (since := "2024-08-09")] alias isSumSq.add := IsSumSq.add -/-- A finite sum of squares is a sum of squares. -/ -theorem IsSumSq.sum_mul_self {ι : Type*} [AddCommMonoid R] (s : Finset ι) (f : ι → R) : - IsSumSq (∑ i ∈ s, f i * f i) := by - induction s using Finset.cons_induction with - | empty => - simpa only [Finset.sum_empty] using IsSumSq.zero - | cons i s his h => - exact (Finset.sum_cons (β := R) his) ▸ IsSumSq.sq_add (f i) (∑ i ∈ s, f i * f i) h +namespace AddSubmonoid +variable {T : Type*} [AddMonoid T] [Mul T] {s : T} -@[deprecated (since := "2024-12-27")] alias isSumSq_sum_mul_self := IsSumSq.sum_mul_self - -variable (R) in +variable (T) in /-- -In an additive monoid with multiplication `R`, the type `sumSqIn R` is the submonoid of sums of +In an additive monoid with multiplication `R`, `AddSubmonoid.sumSq R` is the submonoid of sums of squares in `R`. -/ -def sumSqIn [AddMonoid R] : AddSubmonoid R where - carrier := {S : R | IsSumSq S} - zero_mem' := IsSumSq.zero - add_mem' := IsSumSq.add +@[simps] +def sumSq [AddMonoid T] : AddSubmonoid T where + carrier := {s : T | IsSumSq s} + zero_mem' := .zero + add_mem' := .add + +attribute [norm_cast] coe_sumSq -@[deprecated (since := "2024-08-09")] alias SumSqIn := sumSqIn +@[simp] theorem mem_sumSq : s ∈ sumSq T ↔ IsSumSq s := Iff.rfl + +end AddSubmonoid + +@[deprecated (since := "2024-08-09")] alias SumSqIn := AddSubmonoid.sumSq +@[deprecated (since := "2025-01-03")] alias sumSqIn := AddSubmonoid.sumSq +@[deprecated (since := "2025-01-06")] alias sumSq := AddSubmonoid.sumSq /-- -In an additive monoid with multiplication, every square is a sum of squares. By definition, a square -in `R` is a term `x : R` such that `x = y * y` for some `y : R` and in Mathlib this is known as -`IsSquare R` (see Mathlib.Algebra.Group.Even). +In an additive monoid with multiplication, squares are sums of squares +(see Mathlib.Algebra.Group.Even). -/ -theorem mem_sumSqIn_of_isSquare [AddMonoid R] {x : R} (px : IsSquare x) : x ∈ sumSqIn R := by - rcases px with ⟨y, py⟩ - rw [py, ← AddMonoid.add_zero (y * y)] - exact IsSumSq.sq_add _ _ IsSumSq.zero +theorem mem_sumSq_of_isSquare [AddMonoid R] [Mul R] {x : R} (hx : IsSquare x) : + x ∈ AddSubmonoid.sumSq R := by + rcases hx with ⟨y, hy⟩ + rw [hy, ← AddMonoid.add_zero (y * y)] + exact IsSumSq.sq_add _ IsSumSq.zero -@[deprecated (since := "2024-08-09")] alias SquaresInSumSq := mem_sumSqIn_of_isSquare +@[deprecated (since := "2024-08-09")] alias SquaresInSumSq := mem_sumSq_of_isSquare +@[deprecated (since := "2025-01-03")] alias mem_sumSqIn_of_isSquare := mem_sumSq_of_isSquare /-- In an additive monoid with multiplication `R`, the submonoid generated by the squares is the set of -sums of squares. +sums of squares in `R`. -/ -theorem AddSubmonoid.closure_isSquare [AddMonoid R] : closure {x : R | IsSquare x} = sumSqIn R := by - refine le_antisymm (closure_le.2 (fun x hx ↦ mem_sumSqIn_of_isSquare hx)) (fun x hx ↦ ?_) - induction hx with - | zero => exact zero_mem _ - | sq_add a S _ ih => exact AddSubmonoid.add_mem _ (subset_closure ⟨a, rfl⟩) ih +theorem AddSubmonoid.closure_isSquare [AddMonoid R] [Mul R] : + closure {x : R | IsSquare x} = sumSq R := by + refine le_antisymm (closure_le.2 (fun x hx ↦ mem_sumSq_of_isSquare hx)) (fun x hx ↦ ?_) + induction hx <;> aesop @[deprecated (since := "2024-08-09")] alias SquaresAddClosure := AddSubmonoid.closure_isSquare /-- -Let `R` be a linearly ordered semiring in which the property `a ≤ b → ∃ c, a + c = b` holds -(e.g. `R = ℕ`). If `S : R` is a sum of squares in `R`, then `0 ≤ S`. This is used in -`Mathlib.Algebra.Ring.Semireal.Defs` to show that linearly ordered fields are semireal. +In an additive commutative monoid with multiplication, +`∑ i ∈ I, a i * a i` is a sum of squares. +-/ +theorem IsSumSq.sum_mul_self [AddCommMonoid R] [Mul R] {ι : Type*} (I : Finset ι) (a : ι → R) : + IsSumSq (∑ i ∈ I, a i * a i) := by + induction I using Finset.cons_induction with + | empty => simpa using IsSumSq.zero + | cons i _ hi h => exact (Finset.sum_cons (β := R) hi) ▸ IsSumSq.sq_add (a i) h + +@[deprecated (since := "2024-12-27")] alias isSumSq_sum_mul_self := IsSumSq.sum_mul_self + +/-- +In a linearly ordered semiring with the property `a ≤ b → ∃ c, a + c = b` (e.g. `ℕ`), +sums of squares are non-negative. -/ -theorem IsSumSq.nonneg {R : Type*} [LinearOrderedSemiring R] [ExistsAddOfLE R] {S : R} - (pS : IsSumSq S) : 0 ≤ S := by - induction pS with - | zero => simp only [le_refl] - | sq_add x S _ ih => apply add_nonneg ?_ ih; simp only [← pow_two x, sq_nonneg] +theorem IsSumSq.nonneg {R : Type*} [LinearOrderedSemiring R] [ExistsAddOfLE R] {s : R} + (hs : IsSumSq s) : 0 ≤ s := by + induction hs with + | zero => simp only [le_refl] + | sq_add x _ ih => apply add_nonneg ?_ ih; simp only [← pow_two x, sq_nonneg] @[deprecated (since := "2024-08-09")] alias isSumSq.nonneg := IsSumSq.nonneg diff --git a/Mathlib/Algebra/RingQuot.lean b/Mathlib/Algebra/RingQuot.lean index fb0556615bab7..468b97fb7ad6c 100644 --- a/Mathlib/Algebra/RingQuot.lean +++ b/Mathlib/Algebra/RingQuot.lean @@ -412,7 +412,7 @@ theorem ringQuot_ext [Semiring T] {r : R → R → Prop} (f g : RingQuot r →+* (w : f.comp (mkRingHom r) = g.comp (mkRingHom r)) : f = g := by ext x rcases mkRingHom_surjective r x with ⟨x, rfl⟩ - exact (RingHom.congr_fun w x : _) + exact (RingHom.congr_fun w x :) variable [Semiring T] diff --git a/Mathlib/Algebra/Squarefree/Basic.lean b/Mathlib/Algebra/Squarefree/Basic.lean index bdc6c55ec4a16..6096792244d94 100644 --- a/Mathlib/Algebra/Squarefree/Basic.lean +++ b/Mathlib/Algebra/Squarefree/Basic.lean @@ -160,9 +160,6 @@ theorem Squarefree.isRadical {x : R} (hx : Squarefree x) : IsRadical x := theorem Squarefree.dvd_pow_iff_dvd {x y : R} {n : ℕ} (hsq : Squarefree x) (h0 : n ≠ 0) : x ∣ y ^ n ↔ x ∣ y := ⟨hsq.isRadical n y, (·.pow h0)⟩ -@[deprecated (since := "2024-02-12")] -alias UniqueFactorizationMonoid.dvd_pow_iff_dvd_of_squarefree := Squarefree.dvd_pow_iff_dvd - end variable [CancelCommMonoidWithZero R] {x y p d : R} diff --git a/Mathlib/Algebra/Star/Basic.lean b/Mathlib/Algebra/Star/Basic.lean index 04f2f59e38df2..929ae83b99d4a 100644 --- a/Mathlib/Algebra/Star/Basic.lean +++ b/Mathlib/Algebra/Star/Basic.lean @@ -289,7 +289,7 @@ theorem star_natCast [NonAssocSemiring R] [StarRing R] (n : ℕ) : star (n : R) @[simp] theorem star_ofNat [NonAssocSemiring R] [StarRing R] (n : ℕ) [n.AtLeastTwo] : - star (no_index (OfNat.ofNat n) : R) = OfNat.ofNat n := + star (ofNat(n) : R) = OfNat.ofNat n := star_natCast _ section diff --git a/Mathlib/Algebra/Star/NonUnitalSubalgebra.lean b/Mathlib/Algebra/Star/NonUnitalSubalgebra.lean index cca41ea3fd0c1..7fe2b7b29227b 100644 --- a/Mathlib/Algebra/Star/NonUnitalSubalgebra.lean +++ b/Mathlib/Algebra/Star/NonUnitalSubalgebra.lean @@ -445,7 +445,7 @@ theorem coe_codRestrict (f : F) (S : NonUnitalStarSubalgebra R B) (hf : ∀ x, f theorem injective_codRestrict (f : F) (S : NonUnitalStarSubalgebra R B) (hf : ∀ x : A, f x ∈ S) : Function.Injective (NonUnitalStarAlgHom.codRestrict f S hf) ↔ Function.Injective f := - ⟨fun H _x _y hxy => H <| Subtype.eq hxy, fun H _x _y hxy => H (congr_arg Subtype.val hxy : _)⟩ + ⟨fun H _x _y hxy => H <| Subtype.eq hxy, fun H _x _y hxy => H (congr_arg Subtype.val hxy :)⟩ /-- Restrict the codomain of a non-unital star algebra homomorphism `f` to `f.range`. diff --git a/Mathlib/Algebra/Star/StarAlgHom.lean b/Mathlib/Algebra/Star/StarAlgHom.lean index 3a2986c6d730c..c1f51d104dfa9 100644 --- a/Mathlib/Algebra/Star/StarAlgHom.lean +++ b/Mathlib/Algebra/Star/StarAlgHom.lean @@ -279,7 +279,7 @@ lemma coe_restrictScalars' (f : A →⋆ₙₐ[S] B) : (f.restrictScalars R : A theorem restrictScalars_injective : Function.Injective (restrictScalars R : (A →⋆ₙₐ[S] B) → A →⋆ₙₐ[R] B) := - fun _ _ h ↦ ext (DFunLike.congr_fun h : _) + fun _ _ h ↦ ext (DFunLike.congr_fun h :) end RestrictScalars diff --git a/Mathlib/Algebra/Star/Subalgebra.lean b/Mathlib/Algebra/Star/Subalgebra.lean index 7df842838e470..8761b250d94cb 100644 --- a/Mathlib/Algebra/Star/Subalgebra.lean +++ b/Mathlib/Algebra/Star/Subalgebra.lean @@ -758,7 +758,7 @@ theorem subtype_comp_codRestrict (f : A →⋆ₐ[R] B) (S : StarSubalgebra R B) theorem injective_codRestrict (f : A →⋆ₐ[R] B) (S : StarSubalgebra R B) (hf : ∀ x : A, f x ∈ S) : Function.Injective (StarAlgHom.codRestrict f S hf) ↔ Function.Injective f := - ⟨fun H _x _y hxy => H <| Subtype.eq hxy, fun H _x _y hxy => H (congr_arg Subtype.val hxy : _)⟩ + ⟨fun H _x _y hxy => H <| Subtype.eq hxy, fun H _x _y hxy => H (congr_arg Subtype.val hxy :)⟩ /-- Restriction of the codomain of a `StarAlgHom` to its range. -/ def rangeRestrict (f : A →⋆ₐ[R] B) : A →⋆ₐ[R] f.range := diff --git a/Mathlib/AlgebraicGeometry/Cover/Open.lean b/Mathlib/AlgebraicGeometry/Cover/Open.lean index 7c743cdacbf51..49feb36d291ee 100644 --- a/Mathlib/AlgebraicGeometry/Cover/Open.lean +++ b/Mathlib/AlgebraicGeometry/Cover/Open.lean @@ -278,7 +278,7 @@ theorem affineBasisCover_map_range (X : Scheme.{u}) (x : X) erw [coe_comp, Set.range_comp] -- Porting note: `congr` fails to see the goal is comparing image of the same function refine congr_arg (_ '' ·) ?_ - exact (PrimeSpectrum.localization_away_comap_range (Localization.Away r) r : _) + exact (PrimeSpectrum.localization_away_comap_range (Localization.Away r) r :) theorem affineBasisCover_is_basis (X : Scheme.{u}) : TopologicalSpace.IsTopologicalBasis diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean index 189a1bc957349..97e4ff84b8c19 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean @@ -835,9 +835,7 @@ lemma baseChange_negPolynomial : rw [← map_negPolynomial, map_baseChange] lemma baseChange_negY (x y : A) : - (W.baseChange B).toAffine.negY (f x) (f y) = f ((W.baseChange A).toAffine.negY x y) := by - erw [← map_negY, map_baseChange] - rfl + (W.baseChange B).toAffine.negY (f x) (f y) = f ((W.baseChange A).toAffine.negY x y) := by simp lemma baseChange_addPolynomial (x y L : A) : (W.baseChange B).toAffine.addPolynomial (f x) (f y) (f L) = @@ -847,21 +845,15 @@ lemma baseChange_addPolynomial (x y L : A) : lemma baseChange_addX (x₁ x₂ L : A) : (W.baseChange B).toAffine.addX (f x₁) (f x₂) (f L) = - f ((W.baseChange A).toAffine.addX x₁ x₂ L) := by - erw [← map_addX, map_baseChange] - rfl + f ((W.baseChange A).toAffine.addX x₁ x₂ L) := by simp lemma baseChange_negAddY (x₁ x₂ y₁ L : A) : (W.baseChange B).toAffine.negAddY (f x₁) (f x₂) (f y₁) (f L) = - f ((W.baseChange A).toAffine.negAddY x₁ x₂ y₁ L) := by - erw [← map_negAddY, map_baseChange] - rfl + f ((W.baseChange A).toAffine.negAddY x₁ x₂ y₁ L) := by simp lemma baseChange_addY (x₁ x₂ y₁ L : A) : (W.baseChange B).toAffine.addY (f x₁) (f x₂) (f y₁) (f L) = - f ((W.baseChange A).toAffine.addY x₁ x₂ y₁ L) := by - erw [← map_addY, map_baseChange] - rfl + f ((W.baseChange A).toAffine.addY x₁ x₂ y₁ L) := by simp variable {F : Type u} [Field F] [Algebra R F] [Algebra S F] [IsScalarTower R S F] {K : Type v} [Field K] [Algebra R K] [Algebra S K] [IsScalarTower R S K] (f : F →ₐ[S] K) diff --git a/Mathlib/AlgebraicGeometry/Limits.lean b/Mathlib/AlgebraicGeometry/Limits.lean index b4f7a0106ddca..6ca38222835bb 100644 --- a/Mathlib/AlgebraicGeometry/Limits.lean +++ b/Mathlib/AlgebraicGeometry/Limits.lean @@ -115,7 +115,7 @@ instance : HasInitial Scheme.{u} := hasInitial_of_unique ∅ instance initial_isEmpty : IsEmpty (⊥_ Scheme) := - ⟨fun x => ((initial.to Scheme.empty : _).base x).elim⟩ + ⟨fun x => ((initial.to Scheme.empty :).base x).elim⟩ theorem isAffineOpen_bot (X : Scheme) : IsAffineOpen (⊥ : X.Opens) := @isAffine_of_isEmpty _ (inferInstanceAs (IsEmpty (∅ : Set X))) @@ -276,7 +276,7 @@ lemma disjoint_opensRange_sigmaι (i j : ι) (h : i ≠ j) : obtain ⟨rfl⟩ := (sigmaι_eq_iff _ _ _ _ _).mp hy cases h rfl -lemma exists_sigmaι_eq (x : (∐ f : _)) : ∃ i y, (Sigma.ι f i).base y = x := by +lemma exists_sigmaι_eq (x : (∐ f :)) : ∃ i y, (Sigma.ι f i).base y = x := by obtain ⟨i, y, e⟩ := (disjointGlueData f).ι_jointly_surjective ((sigmaIsoGlued f).hom.base x) refine ⟨i, y, (sigmaIsoGlued f).hom.isOpenEmbedding.injective ?_⟩ rwa [← Scheme.comp_base_apply, ι_sigmaIsoGlued_hom] @@ -296,7 +296,7 @@ def sigmaOpenCover : (∐ f).OpenCover where /-- The underlying topological space of the coproduct is homeomorphic to the disjoint union. -/ noncomputable -def sigmaMk : (Σ i, f i) ≃ₜ (∐ f : _) := +def sigmaMk : (Σ i, f i) ≃ₜ (∐ f :) := TopCat.homeoOfIso ((colimit.isoColimitCocone ⟨_, TopCat.sigmaCofanIsColimit _⟩).symm ≪≫ (PreservesCoproduct.iso Scheme.forgetToTop f).symm) diff --git a/Mathlib/AlgebraicGeometry/Modules/Tilde.lean b/Mathlib/AlgebraicGeometry/Modules/Tilde.lean index affc194a86b50..aa13874400ecd 100644 --- a/Mathlib/AlgebraicGeometry/Modules/Tilde.lean +++ b/Mathlib/AlgebraicGeometry/Modules/Tilde.lean @@ -182,7 +182,7 @@ namespace Tilde @[simp] theorem res_apply (U V : Opens (PrimeSpectrum.Top R)) (i : V ⟶ U) (s : (tildeInModuleCat M).obj (op U)) (x : V) : - ((tildeInModuleCat M).map i.op s).1 x = (s.1 (i x) : _) := + ((tildeInModuleCat M).map i.op s).1 x = (s.1 (i x) :) := rfl lemma smul_section_apply (r : R) (U : Opens (PrimeSpectrum.Top R)) @@ -273,7 +273,7 @@ def openToLocalization (U : Opens (PrimeSpectrum R)) (x : PrimeSpectrum R) (hx : ModuleCat.ofHom (X := (tildeInModuleCat M).obj (op U)) (Y := ModuleCat.of R (LocalizedModule x.asIdeal.primeCompl M)) - { toFun := fun s => (s.1 ⟨x, hx⟩ : _) + { toFun := fun s => (s.1 ⟨x, hx⟩ :) map_add' := fun _ _ => rfl map_smul' := fun _ _ => rfl } @@ -299,7 +299,7 @@ theorem germ_comp_stalkToFiberLinearMap (U : Opens (PrimeSpectrum.Top R)) (x) (h theorem stalkToFiberLinearMap_germ (U : Opens (PrimeSpectrum.Top R)) (x : PrimeSpectrum.Top R) (hx : x ∈ U) (s : (tildeInModuleCat M).1.obj (op U)) : (stalkToFiberLinearMap M x).hom - (TopCat.Presheaf.germ (tildeInModuleCat M) U x hx s) = (s.1 ⟨x, hx⟩ : _) := + (TopCat.Presheaf.germ (tildeInModuleCat M) U x hx s) = (s.1 ⟨x, hx⟩ :) := DFunLike.ext_iff.1 (ModuleCat.hom_ext_iff.mp (germ_comp_stalkToFiberLinearMap M U x hx)) s @[reassoc (attr := simp), elementwise (attr := simp)] diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Affine.lean b/Mathlib/AlgebraicGeometry/Morphisms/Affine.lean index 32767133cd5ef..ff4190d85fb10 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Affine.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Affine.lean @@ -84,8 +84,7 @@ instance {X : Scheme} (r : Γ(X, ⊤)) : rw [X.basicOpen_res] ext1 refine Set.image_preimage_eq_inter_range.trans ?_ - erw [Subtype.range_coe] - rfl + simp lemma isAffineOpen_of_isAffineOpen_basicOpen_aux (s : Set Γ(X, ⊤)) (hs : Ideal.span s = ⊤) (hs₂ : ∀ i ∈ s, IsAffineOpen (X.basicOpen i)) : diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Preimmersion.lean b/Mathlib/AlgebraicGeometry/Morphisms/Preimmersion.lean index c1e4acd9220c2..b930b47ba6a4c 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Preimmersion.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Preimmersion.lean @@ -116,7 +116,7 @@ instance : IsStableUnderBaseChange @IsPreimmersion := by refine .mk' fun X Y Z f g _ _ ↦ ?_ have := pullback_fst (P := @SurjectiveOnStalks) f g inferInstance constructor - let L (x : (pullback f g : _)) : { x : X × Y | f.base x.1 = g.base x.2 } := + let L (x : (pullback f g :)) : { x : X × Y | f.base x.1 = g.base x.2 } := ⟨⟨(pullback.fst f g).base x, (pullback.snd f g).base x⟩, by simp only [Set.mem_setOf, ← Scheme.comp_base_apply, pullback.condition]⟩ have : IsEmbedding L := IsEmbedding.of_comp (by fun_prop) continuous_subtype_val diff --git a/Mathlib/AlgebraicGeometry/Morphisms/SurjectiveOnStalks.lean b/Mathlib/AlgebraicGeometry/Morphisms/SurjectiveOnStalks.lean index b71393e088b97..79e2e54a20a2b 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/SurjectiveOnStalks.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/SurjectiveOnStalks.lean @@ -129,7 +129,7 @@ lemma isEmbedding_pullback {X Y S : Scheme.{u}} (f : X ⟶ S) (g : Y ⟶ S) [Sur (continuous_fst.1 _ ((𝒱 ijk.1).map ijk.2.1 ≫ (𝒰.pullbackCover f).map ijk.1).opensRange.2).inter (continuous_snd.1 _ ((𝒲 ijk.1).map ijk.2.2 ≫ (𝒰.pullbackCover g).map ijk.1).opensRange.2)⟩ - have : Set.range L ⊆ (iSup U : _) := by + have : Set.range L ⊆ (iSup U :) := by simp only [Scheme.Cover.pullbackCover_J, Scheme.Cover.pullbackCover_obj, Set.range_subset_iff] intro z simp only [SetLike.mem_coe, TopologicalSpace.Opens.mem_iSup, Sigma.exists, Prod.exists] diff --git a/Mathlib/AlgebraicGeometry/Morphisms/UniversallyClosed.lean b/Mathlib/AlgebraicGeometry/Morphisms/UniversallyClosed.lean index 32c4886d35bf2..d04850a4a2454 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/UniversallyClosed.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/UniversallyClosed.lean @@ -109,7 +109,7 @@ lemma compactSpace_of_universallyClosed let Ti (i : 𝒰.J) : T.Opens := basicOpen (MvPolynomial.X i) let fT : pullback f q ⟶ T := pullback.snd f q let p : pullback f q ⟶ X := pullback.fst f q - let Z : Set (pullback f q : _) := (⨆ i, fT ⁻¹ᵁ (Ti i) ⊓ p ⁻¹ᵁ (U i) : (pullback f q).Opens)ᶜ + let Z : Set (pullback f q :) := (⨆ i, fT ⁻¹ᵁ (Ti i) ⊓ p ⁻¹ᵁ (U i) : (pullback f q).Opens)ᶜ have hZ : IsClosed Z := by simp only [Z, isClosed_compl_iff, Opens.coe_iSup, Opens.coe_inf, Opens.map_coe] exact isOpen_iUnion fun i ↦ (fT.continuous.1 _ (Ti i).2).inter (p.continuous.1 _ (U i).2) diff --git a/Mathlib/AlgebraicGeometry/OpenImmersion.lean b/Mathlib/AlgebraicGeometry/OpenImmersion.lean index 0c8626614a15c..754e130626539 100644 --- a/Mathlib/AlgebraicGeometry/OpenImmersion.lean +++ b/Mathlib/AlgebraicGeometry/OpenImmersion.lean @@ -46,7 +46,7 @@ protected def scheme (X : LocallyRingedSpace.{u}) (h : ∀ x : X, ∃ (R : CommRingCat) (f : Spec.toLocallyRingedSpace.obj (op R) ⟶ X), - (x ∈ Set.range f.base : _) ∧ LocallyRingedSpace.IsOpenImmersion f) : + (x ∈ Set.range f.base :) ∧ LocallyRingedSpace.IsOpenImmersion f) : Scheme where toLocallyRingedSpace := X local_affine := by @@ -65,9 +65,6 @@ theorem IsOpenImmersion.isOpen_range {X Y : Scheme.{u}} (f : X ⟶ Y) [H : IsOpe IsOpen (Set.range f.base) := H.base_open.isOpen_range -@[deprecated (since := "2024-03-17")] -alias IsOpenImmersion.open_range := IsOpenImmersion.isOpen_range - namespace Scheme.Hom variable {X Y : Scheme.{u}} (f : Scheme.Hom X Y) [H : IsOpenImmersion f] @@ -225,7 +222,7 @@ namespace Scheme instance basic_open_isOpenImmersion {R : CommRingCat.{u}} (f : R) : IsOpenImmersion (Spec.map (CommRingCat.ofHom (algebraMap R (Localization.Away f)))) := by apply SheafedSpace.IsOpenImmersion.of_stalk_iso (H := ?_) - · exact (PrimeSpectrum.localization_away_isOpenEmbedding (Localization.Away f) f : _) + · exact (PrimeSpectrum.localization_away_isOpenEmbedding (Localization.Away f) f :) · intro x exact Spec_map_localization_isIso R (Submonoid.powers f) x @@ -416,7 +413,7 @@ theorem _root_.AlgebraicGeometry.isIso_iff_stalk_iso {X Y : Scheme.{u}} (f : X · intro H; exact ⟨inferInstance, (TopCat.homeoOfIso (asIso f.base)).isOpenEmbedding⟩ /-- An open immersion induces an isomorphism from the domain onto the image -/ -def isoRestrict : X ≅ (Z.restrict H.base_open : _) := +def isoRestrict : X ≅ (Z.restrict H.base_open :) := Scheme.fullyFaithfulForgetToLocallyRingedSpace.preimageIso (LocallyRingedSpace.IsOpenImmersion.isoRestrict f.toLRSHom) diff --git a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean index 680ed47f8dce6..5111ffb0d516b 100644 --- a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean +++ b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean @@ -142,7 +142,7 @@ theorem t1Space_iff_isField [IsDomain R] : T1Space (PrimeSpectrum R) ↔ IsField (mt (Ring.ne_bot_of_isMaximal_of_not_isField <| (isClosed_singleton_iff_isMaximal _).1 (T1Space.t1 ⟨⊥, hbot⟩)) - (by aesop)) + (by simp)) · refine ⟨fun x => (isClosed_singleton_iff_isMaximal x).2 ?_⟩ by_cases hx : x.asIdeal = ⊥ · letI := h.toSemifield @@ -649,7 +649,7 @@ def localizationMapOfSpecializes {x y : PrimeSpectrum R} (h : x ⤳ y) : rw [← PrimeSpectrum.le_iff_specializes, ← asIdeal_le_asIdeal, ← SetLike.coe_subset_coe, ← Set.compl_subset_compl] at h exact (IsLocalization.map_units (Localization.AtPrime x.asIdeal) - ⟨a, show a ∈ x.asIdeal.primeCompl from h ha⟩ : _)) + ⟨a, show a ∈ x.asIdeal.primeCompl from h ha⟩ :)) section stableUnderSpecialization diff --git a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean index ba5a97a21ee36..91892e6714502 100644 --- a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean +++ b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean @@ -491,7 +491,6 @@ lemma toSpec_fromSpec {f : A} {m : ℕ} (f_deg : f ∈ 𝒜 m) (hm : 0 < m) (x : rw [← FromSpec.num_mem_carrier_iff f_deg hm x] exact ToSpec.mk_mem_carrier _ z -@[deprecated (since := "2024-03-02")] alias toSpecFromSpec := toSpec_fromSpec end toSpecFromSpec diff --git a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/StructureSheaf.lean b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/StructureSheaf.lean index 7532e95080f90..44912ce1425be 100644 --- a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/StructureSheaf.lean +++ b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/StructureSheaf.lean @@ -243,7 +243,7 @@ evaluates the section on the point corresponding to a given homogeneous prime id def openToLocalization (U : Opens (ProjectiveSpectrum.top 𝒜)) (x : ProjectiveSpectrum.top 𝒜) (hx : x ∈ U) : (Proj.structureSheaf 𝒜).1.obj (op U) ⟶ CommRingCat.of (at x) := CommRingCat.ofHom - { toFun s := (s.1 ⟨x, hx⟩ : _) + { toFun s := (s.1 ⟨x, hx⟩ :) map_one' := rfl map_mul' _ _ := rfl map_zero' := rfl diff --git a/Mathlib/AlgebraicGeometry/Restrict.lean b/Mathlib/AlgebraicGeometry/Restrict.lean index afcdc6496a7fb..766c1a26259da 100644 --- a/Mathlib/AlgebraicGeometry/Restrict.lean +++ b/Mathlib/AlgebraicGeometry/Restrict.lean @@ -301,7 +301,7 @@ isomorphic to the structure sheaf. -/ @[simps!] def Scheme.restrictFunctorΓ : X.restrictFunctor.op ⋙ (Over.forget X).op ⋙ Scheme.Γ ≅ X.presheaf := NatIso.ofComponents - (fun U => X.presheaf.mapIso ((eqToIso (unop U).isOpenEmbedding_obj_top).symm.op : _)) + (fun U => X.presheaf.mapIso ((eqToIso (unop U).isOpenEmbedding_obj_top).symm.op :)) (by intro U V i dsimp diff --git a/Mathlib/AlgebraicGeometry/Scheme.lean b/Mathlib/AlgebraicGeometry/Scheme.lean index 5fffbe52e634b..6acbeea9a53e1 100644 --- a/Mathlib/AlgebraicGeometry/Scheme.lean +++ b/Mathlib/AlgebraicGeometry/Scheme.lean @@ -248,16 +248,16 @@ unif_hint forgetToTop_obj_eq_coe (X : Scheme) where ⊢ forgetToTop.obj X ≟ (X : TopCat) @[simp] -theorem id.base (X : Scheme) : (𝟙 X : _).base = 𝟙 _ := +theorem id.base (X : Scheme) : (𝟙 X :).base = 𝟙 _ := rfl @[simp] theorem id_app {X : Scheme} (U : X.Opens) : - (𝟙 X : _).app U = 𝟙 _ := rfl + (𝟙 X :).app U = 𝟙 _ := rfl @[simp] theorem id_appTop {X : Scheme} : - (𝟙 X : _).appTop = 𝟙 _ := + (𝟙 X :).appTop = 𝟙 _ := rfl @[reassoc] diff --git a/Mathlib/AlgebraicGeometry/StructureSheaf.lean b/Mathlib/AlgebraicGeometry/StructureSheaf.lean index 8d35ebe6a7520..e6e5296cb7084 100644 --- a/Mathlib/AlgebraicGeometry/StructureSheaf.lean +++ b/Mathlib/AlgebraicGeometry/StructureSheaf.lean @@ -263,7 +263,7 @@ namespace StructureSheaf @[simp] theorem res_apply (U V : Opens (PrimeSpectrum.Top R)) (i : V ⟶ U) (s : (structureSheaf R).1.obj (op U)) (x : V) : - ((structureSheaf R).1.map i.op s).1 x = (s.1 (i x) : _) := + ((structureSheaf R).1.map i.op s).1 x = (s.1 (i x) :) := rfl /- @@ -460,7 +460,7 @@ the section on the point corresponding to a given prime ideal. -/ def openToLocalization (U : Opens (PrimeSpectrum.Top R)) (x : PrimeSpectrum.Top R) (hx : x ∈ U) : (structureSheaf R).1.obj (op U) ⟶ CommRingCat.of (Localization.AtPrime x.asIdeal) := CommRingCat.ofHom - { toFun s := (s.1 ⟨x, hx⟩ : _) + { toFun s := (s.1 ⟨x, hx⟩ :) map_one' := rfl map_mul' _ _ := rfl map_zero' := rfl @@ -471,12 +471,12 @@ theorem coe_openToLocalization (U : Opens (PrimeSpectrum.Top R)) (x : PrimeSpect (hx : x ∈ U) : (openToLocalization R U x hx : (structureSheaf R).1.obj (op U) → Localization.AtPrime x.asIdeal) = - fun s => (s.1 ⟨x, hx⟩ : _) := + fun s => s.1 ⟨x, hx⟩ := rfl theorem openToLocalization_apply (U : Opens (PrimeSpectrum.Top R)) (x : PrimeSpectrum.Top R) (hx : x ∈ U) (s : (structureSheaf R).1.obj (op U)) : - openToLocalization R U x hx s = (s.1 ⟨x, hx⟩ : _) := + openToLocalization R U x hx s = s.1 ⟨x, hx⟩ := rfl /-- The ring homomorphism from the stalk of the structure sheaf of `R` at a point corresponding to @@ -650,7 +650,7 @@ theorem locally_const_basicOpen (U : Opens (PrimeSpectrum.Top R)) have basic_opens_eq := PrimeSpectrum.basicOpen_pow h (n + 1) (by omega) have i_basic_open := eqToHom basic_opens_eq ≫ homOfLE hDhV -- We claim that `(f * c) / h ^ (n+1)` is our desired representation - use f * c, h ^ (n + 1), i_basic_open ≫ iVU, (basic_opens_eq.symm.le : _) hxDh + use f * c, h ^ (n + 1), i_basic_open ≫ iVU, (basic_opens_eq.symm.le :) hxDh rw [op_comp, Functor.map_comp] --, comp_apply, ← s_eq, res_const] -- Porting note: `comp_apply` can't be rewritten, so use a change change const R _ _ _ _ = (structureSheaf R).1.map i_basic_open.op @@ -661,7 +661,7 @@ theorem locally_const_basicOpen (U : Opens (PrimeSpectrum.Top R)) swap · intro y hy rw [basic_opens_eq] at hy - exact (Set.Subset.trans hDhV hVDg : _) hy + exact (Set.Subset.trans hDhV hVDg :) hy -- All that is left is a simple calculation apply const_ext rw [mul_assoc f c g, hc] @@ -842,7 +842,7 @@ theorem toBasicOpen_surjective (f : R) : Function.Surjective (toBasicOpen R f) : · intro y hy change y ∈ PrimeSpectrum.basicOpen (f ^ (n + 1)) rw [PrimeSpectrum.basicOpen_pow f (n + 1) (by omega)] - exact (leOfHom (iDh i) : _) hy + exact (leOfHom (iDh i) :) hy -- The rest of the proof is just computation apply const_ext rw [← hb, Finset.sum_mul, Finset.mul_sum] @@ -989,7 +989,7 @@ def comapFun (f : R →+* S) (U : Opens (PrimeSpectrum.Top R)) (V : Opens (Prime (hUV : V.1 ⊆ PrimeSpectrum.comap f ⁻¹' U.1) (s : ∀ x : U, Localizations R x) (y : V) : Localizations S y := Localization.localRingHom (PrimeSpectrum.comap f y.1).asIdeal _ f rfl - (s ⟨PrimeSpectrum.comap f y.1, hUV y.2⟩ : _) + (s ⟨PrimeSpectrum.comap f y.1, hUV y.2⟩ :) theorem comapFunIsLocallyFraction (f : R →+* S) (U : Opens (PrimeSpectrum.Top R)) (V : Opens (PrimeSpectrum.Top S)) (hUV : V.1 ⊆ PrimeSpectrum.comap f ⁻¹' U.1) @@ -1052,7 +1052,7 @@ theorem comap_apply (f : R →+* S) (U : Opens (PrimeSpectrum.Top R)) (s : (structureSheaf R).1.obj (op U)) (p : V) : (comap f U V hUV s).1 p = Localization.localRingHom (PrimeSpectrum.comap f p.1).asIdeal _ f rfl - (s.1 ⟨PrimeSpectrum.comap f p.1, hUV p.2⟩ : _) := + (s.1 ⟨PrimeSpectrum.comap f p.1, hUV p.2⟩ :) := rfl theorem comap_const (f : R →+* S) (U : Opens (PrimeSpectrum.Top R)) diff --git a/Mathlib/AlgebraicTopology/CechNerve.lean b/Mathlib/AlgebraicTopology/CechNerve.lean index 667421531fc21..01886468c2e25 100644 --- a/Mathlib/AlgebraicTopology/CechNerve.lean +++ b/Mathlib/AlgebraicTopology/CechNerve.lean @@ -50,7 +50,7 @@ variable [∀ n : ℕ, HasWidePullback.{0} f.right (fun _ : Fin (n + 1) => f.lef def cechNerve : SimplicialObject C where obj n := widePullback.{0} f.right (fun _ : Fin (n.unop.len + 1) => f.left) fun _ => f.hom map g := WidePullback.lift (WidePullback.base _) - (fun i => WidePullback.π _ (g.unop.toOrderHom i)) (by aesop_cat) + (fun i => WidePullback.π _ (g.unop.toOrderHom i)) (by simp) /-- The morphism between Čech nerves associated to a morphism of arrows. -/ @[simps] @@ -117,10 +117,7 @@ def equivalenceLeftToRight (X : SimplicialObject.Augmented C) (F : Arrow C) left := { app := fun x => Limits.WidePullback.lift (X.hom.app _ ≫ G.right) - (fun i => X.left.map (SimplexCategory.const _ x.unop i).op ≫ G.left) fun i => by - dsimp - erw [Category.assoc, Arrow.w, Augmented.toArrow_obj_hom, NatTrans.naturality_assoc, - Functor.const_obj_map, Category.id_comp] + (fun i => X.left.map (SimplexCategory.const _ x.unop i).op ≫ G.left) fun i => by simp naturality := by intro x y f dsimp @@ -248,7 +245,7 @@ def augmentedCechConerve : Arrow C ⥤ CosimplicialObject.Augmented C where def equivalenceLeftToRight (F : Arrow C) (X : CosimplicialObject.Augmented C) (G : F.augmentedCechConerve ⟶ X) : F ⟶ Augmented.toArrow.obj X where left := G.left - right := (WidePushout.ι _ 0 ≫ G.right.app (SimplexCategory.mk 0) : _) + right := (WidePushout.ι _ 0 ≫ G.right.app (SimplexCategory.mk 0) :) w := by dsimp rw [@WidePushout.arrow_ι_assoc _ _ _ _ _ (fun (_ : Fin 1) => F.hom) @@ -280,9 +277,7 @@ def equivalenceRightToLeft (F : Arrow C) (X : CosimplicialObject.Augmented C) rw [Category.assoc, ← X.right.map_comp] rfl · dsimp - simp only [Functor.const_obj_map, ← NatTrans.naturality, WidePushout.head_desc_assoc, - WidePushout.head_desc, Category.assoc] - erw [Category.id_comp] } + simp [← NatTrans.naturality] } /-- A helper function used in defining the Čech conerve adjunction. -/ @[simps] @@ -408,10 +403,7 @@ def iso (X : C) : (Arrow.mk (terminal.from X)).cechNerve ≅ cechNerveTerminalFr NatIso.ofComponents (fun _ => wideCospan.limitIsoPi _ _) (fun {m n} f => by dsimp only [cechNerveTerminalFrom, Arrow.cechNerve] ext ⟨j⟩ - simp only [Category.assoc, limit.lift_π, Fan.mk_π_app] - erw [wideCospan.limitIsoPi_hom_comp_pi, - wideCospan.limitIsoPi_hom_comp_pi, limit.lift_π] - rfl) + simp) end CechNerveTerminalFrom diff --git a/Mathlib/AlgebraicTopology/DoldKan/FunctorGamma.lean b/Mathlib/AlgebraicTopology/DoldKan/FunctorGamma.lean index 67d99a2628ff5..95e605ee3c7a9 100644 --- a/Mathlib/AlgebraicTopology/DoldKan/FunctorGamma.lean +++ b/Mathlib/AlgebraicTopology/DoldKan/FunctorGamma.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Joël Riou. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Joël Riou -/ -import Mathlib.AlgebraicTopology.SplitSimplicialObject +import Mathlib.AlgebraicTopology.SimplicialObject.Split import Mathlib.AlgebraicTopology.DoldKan.PInfty /-! diff --git a/Mathlib/AlgebraicTopology/DoldKan/Homotopies.lean b/Mathlib/AlgebraicTopology/DoldKan/Homotopies.lean index 1c3db5df719f3..6ecd1804b9639 100644 --- a/Mathlib/AlgebraicTopology/DoldKan/Homotopies.lean +++ b/Mathlib/AlgebraicTopology/DoldKan/Homotopies.lean @@ -145,9 +145,7 @@ theorem hσ'_naturality (q : ℕ) (n m : ℕ) (hnm : c.Rel m n) {X Y : Simplicia unfold hσ split_ifs · rw [zero_comp, comp_zero] - · simp only [zsmul_comp, comp_zsmul] - erw [f.naturality] - rfl + · simp /-- For each q, `Hσ q` is a natural transformation. -/ def natTransHσ (q : ℕ) : alternatingFaceMapComplex C ⟶ alternatingFaceMapComplex C where diff --git a/Mathlib/AlgebraicTopology/DoldKan/Normalized.lean b/Mathlib/AlgebraicTopology/DoldKan/Normalized.lean index cff0ffb6de103..d96dcebf3d5bd 100644 --- a/Mathlib/AlgebraicTopology/DoldKan/Normalized.lean +++ b/Mathlib/AlgebraicTopology/DoldKan/Normalized.lean @@ -111,26 +111,9 @@ using `PInfty` identifies to the composition of the normalized Moore complex fun and the inclusion in the Karoubi envelope. -/ def N₁_iso_normalizedMooreComplex_comp_toKaroubi : N₁ ≅ normalizedMooreComplex A ⋙ toKaroubi _ where hom := - { app := fun X => - { f := PInftyToNormalizedMooreComplex X - comm := by erw [comp_id, PInfty_comp_PInftyToNormalizedMooreComplex] } - naturality := fun X Y f => by - simp only [Functor.comp_map, normalizedMooreComplex_map, - PInftyToNormalizedMooreComplex_naturality, Karoubi.hom_ext_iff, Karoubi.comp_f, N₁_map_f, - PInfty_comp_PInftyToNormalizedMooreComplex_assoc, toKaroubi_map_f, assoc] } + { app := fun X => { f := PInftyToNormalizedMooreComplex X } } inv := - { app := fun X => - { f := inclusionOfMooreComplexMap X - comm := by erw [inclusionOfMooreComplexMap_comp_PInfty, id_comp] } - naturality := fun X Y f => by - ext - simp only [Functor.comp_obj, normalizedMooreComplex_obj, toKaroubi_obj_X, - NormalizedMooreComplex.obj_X, N₁_obj_X, AlternatingFaceMapComplex.obj_X, Functor.comp_map, - normalizedMooreComplex_map, Karoubi.comp_f, toKaroubi_map_f, HomologicalComplex.comp_f, - NormalizedMooreComplex.map_f, inclusionOfMooreComplexMap_f, NormalizedMooreComplex.objX, - factorThru_arrow, N₁_map_f, inclusionOfMooreComplexMap_comp_PInfty_assoc, - AlternatingFaceMapComplex.map_f] - } + { app := fun X => { f := inclusionOfMooreComplexMap X } } hom_inv_id := by ext X : 3 simp only [PInftyToNormalizedMooreComplex_comp_inclusionOfMooreComplexMap, diff --git a/Mathlib/AlgebraicTopology/DoldKan/SplitSimplicialObject.lean b/Mathlib/AlgebraicTopology/DoldKan/SplitSimplicialObject.lean index 1a0800fec44ea..c04554d7644e9 100644 --- a/Mathlib/AlgebraicTopology/DoldKan/SplitSimplicialObject.lean +++ b/Mathlib/AlgebraicTopology/DoldKan/SplitSimplicialObject.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Joël Riou. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Joël Riou -/ -import Mathlib.AlgebraicTopology.SplitSimplicialObject +import Mathlib.AlgebraicTopology.SimplicialObject.Split import Mathlib.AlgebraicTopology.DoldKan.Degeneracies import Mathlib.AlgebraicTopology.DoldKan.FunctorN @@ -154,7 +154,7 @@ noncomputable def nondegComplex : ChainComplex C ℕ where simp only [d, assoc] have eq : K[X].d i j ≫ 𝟙 (X.obj (op [j])) ≫ K[X].d j k ≫ s.πSummand (IndexSet.id (op [k])) = 0 := by - erw [id_comp, HomologicalComplex.d_comp_d_assoc, zero_comp] + simp rw [s.decomposition_id] at eq classical rw [Fintype.sum_eq_add_sum_compl (IndexSet.id (op [j])), add_comp, comp_add, assoc, diff --git a/Mathlib/AlgebraicTopology/SimplicialObject/Basic.lean b/Mathlib/AlgebraicTopology/SimplicialObject/Basic.lean index d8a661f4056b0..ae05211a3fe1e 100644 --- a/Mathlib/AlgebraicTopology/SimplicialObject/Basic.lean +++ b/Mathlib/AlgebraicTopology/SimplicialObject/Basic.lean @@ -834,7 +834,7 @@ object and back is isomorphic to the given object. -/ @[simps!] def SimplicialObject.Augmented.rightOpLeftOpIso (X : SimplicialObject.Augmented C) : X.rightOp.leftOp ≅ X := - Comma.isoMk X.left.rightOpLeftOpIso (CategoryTheory.eqToIso <| by aesop_cat) + Comma.isoMk X.left.rightOpLeftOpIso (CategoryTheory.eqToIso <| by simp) /-- Converting an augmented cosimplicial object to an augmented simplicial object and back is isomorphic to the given object. -/ diff --git a/Mathlib/AlgebraicTopology/SplitSimplicialObject.lean b/Mathlib/AlgebraicTopology/SimplicialObject/Split.lean similarity index 100% rename from Mathlib/AlgebraicTopology/SplitSimplicialObject.lean rename to Mathlib/AlgebraicTopology/SimplicialObject/Split.lean diff --git a/Mathlib/Analysis/Analytic/Basic.lean b/Mathlib/Analysis/Analytic/Basic.lean index af1cd9c970c9d..b89802058d494 100644 --- a/Mathlib/Analysis/Analytic/Basic.lean +++ b/Mathlib/Analysis/Analytic/Basic.lean @@ -327,6 +327,21 @@ theorem min_radius_le_radius_add (p q : FormalMultilinearSeries 𝕜 E F) : theorem radius_neg (p : FormalMultilinearSeries 𝕜 E F) : (-p).radius = p.radius := by simp only [radius, neg_apply, norm_neg] +theorem radius_le_smul {p : FormalMultilinearSeries 𝕜 E F} {c : 𝕜} : p.radius ≤ (c • p).radius := by + simp only [radius, smul_apply] + refine iSup_mono fun r ↦ iSup_mono' fun C ↦ ⟨‖c‖ * C, iSup_mono' fun h ↦ ?_⟩ + simp only [le_refl, exists_prop, and_true] + intro n + rw [norm_smul c (p n), mul_assoc] + gcongr + exact h n + +theorem radius_smul_eq (p : FormalMultilinearSeries 𝕜 E F) {c : 𝕜} + (hc : c ≠ 0) : (c • p).radius = p.radius := by + apply eq_of_le_of_le _ radius_le_smul + conv => rhs; rw [show p = c⁻¹ • (c • p) by simp [smul_smul, inv_mul_cancel₀ hc]] + apply radius_le_smul + @[simp] theorem radius_shift (p : FormalMultilinearSeries 𝕜 E F) : p.shift.radius = p.radius := by simp only [radius, shift, Nat.succ_eq_add_one, ContinuousMultilinearMap.curryRight_norm] diff --git a/Mathlib/Analysis/Analytic/ChangeOrigin.lean b/Mathlib/Analysis/Analytic/ChangeOrigin.lean index 176f6f15e6afc..0e2e48a529be8 100644 --- a/Mathlib/Analysis/Analytic/ChangeOrigin.lean +++ b/Mathlib/Analysis/Analytic/ChangeOrigin.lean @@ -140,9 +140,7 @@ def changeOriginIndexEquiv : (⟨k', l', ⟨s.map (finCongr hkl).toEmbedding, hs'⟩⟩ : Σk l : ℕ, { s : Finset (Fin (k + l)) // s.card = l }) = ⟨k, l, ⟨s, hs⟩⟩ by apply this <;> simp only [hs, add_tsub_cancel_right] - rintro _ _ rfl rfl hkl hs' - simp only [Equiv.refl_toEmbedding, finCongr_refl, Finset.map_refl, eq_self_iff_true, - OrderIso.refl_toEquiv, and_self_iff, heq_iff_eq] + simp right_inv := by rintro ⟨n, s⟩ simp [tsub_add_cancel_of_le (card_finset_fin_le s), finCongr_eq_equivCast] diff --git a/Mathlib/Analysis/Analytic/Composition.lean b/Mathlib/Analysis/Analytic/Composition.lean index 8badf38fc79a0..f7188918787a3 100644 --- a/Mathlib/Analysis/Analytic/Composition.lean +++ b/Mathlib/Analysis/Analytic/Composition.lean @@ -534,7 +534,7 @@ giving the main statement in `comp_partialSum`. -/ power series. See also `comp_partialSum`. -/ def compPartialSumSource (m M N : ℕ) : Finset (Σ n, Fin n → ℕ) := - Finset.sigma (Finset.Ico m M) (fun n : ℕ => Fintype.piFinset fun _i : Fin n => Finset.Ico 1 N : _) + Finset.sigma (Finset.Ico m M) (fun n : ℕ => Fintype.piFinset fun _i : Fin n => Finset.Ico 1 N :) @[simp] theorem mem_compPartialSumSource_iff (m M N : ℕ) (i : Σ n, Fin n → ℕ) : diff --git a/Mathlib/Analysis/Analytic/Constructions.lean b/Mathlib/Analysis/Analytic/Constructions.lean index a1a029bcd24a3..f56a2644814f0 100644 --- a/Mathlib/Analysis/Analytic/Constructions.lean +++ b/Mathlib/Analysis/Analytic/Constructions.lean @@ -62,12 +62,13 @@ theorem analyticOn_const {v : F} {s : Set E} : AnalyticOn 𝕜 (fun _ => v) s := alias analyticWithinOn_const := analyticOn_const /-! -### Addition, negation, subtraction +### Addition, negation, subtraction, scalar multiplication -/ section variable {f g : E → F} {pf pg : FormalMultilinearSeries 𝕜 E F} {s : Set E} {x : E} {r : ℝ≥0∞} + {c : 𝕜} theorem HasFPowerSeriesWithinOnBall.add (hf : HasFPowerSeriesWithinOnBall f pf s x r) (hg : HasFPowerSeriesWithinOnBall g pg s x r) : @@ -163,6 +164,37 @@ theorem AnalyticAt.sub (hf : AnalyticAt 𝕜 f x) (hg : AnalyticAt 𝕜 g x) : AnalyticAt 𝕜 (f - g) x := by simpa only [sub_eq_add_neg] using hf.add hg.neg +theorem HasFPowerSeriesWithinOnBall.const_smul (hf : HasFPowerSeriesWithinOnBall f pf s x r) : + HasFPowerSeriesWithinOnBall (c • f) (c • pf) s x r where + r_le := le_trans hf.r_le pf.radius_le_smul + r_pos := hf.r_pos + hasSum := fun hy h'y => (hf.hasSum hy h'y).const_smul _ + +theorem HasFPowerSeriesOnBall.const_smul (hf : HasFPowerSeriesOnBall f pf x r) : + HasFPowerSeriesOnBall (c • f) (c • pf) x r where + r_le := le_trans hf.r_le pf.radius_le_smul + r_pos := hf.r_pos + hasSum := fun hy => (hf.hasSum hy).const_smul _ + +theorem HasFPowerSeriesWithinAt.const_smul (hf : HasFPowerSeriesWithinAt f pf s x) : + HasFPowerSeriesWithinAt (c • f) (c • pf) s x := + let ⟨_, hrf⟩ := hf + hrf.const_smul.hasFPowerSeriesWithinAt + +theorem HasFPowerSeriesAt.const_smul (hf : HasFPowerSeriesAt f pf x) : + HasFPowerSeriesAt (c • f) (c • pf) x := + let ⟨_, hrf⟩ := hf + hrf.const_smul.hasFPowerSeriesAt + +theorem AnalyticWithinAt.const_smul (hf : AnalyticWithinAt 𝕜 f s x) : + AnalyticWithinAt 𝕜 (c • f) s x := + let ⟨_, hpf⟩ := hf + hpf.const_smul.analyticWithinAt + +theorem AnalyticAt.const_smul (hf : AnalyticAt 𝕜 f x) : AnalyticAt 𝕜 (c • f) x := + let ⟨_, hpf⟩ := hf + hpf.const_smul.analyticAt + theorem AnalyticOn.add (hf : AnalyticOn 𝕜 f s) (hg : AnalyticOn 𝕜 g s) : AnalyticOn 𝕜 (f + g) s := fun z hz => (hf z hz).add (hg z hz) diff --git a/Mathlib/Analysis/BoxIntegral/Partition/SubboxInduction.lean b/Mathlib/Analysis/BoxIntegral/Partition/SubboxInduction.lean index 58f1cc837f67c..ac9965aaeb450 100644 --- a/Mathlib/Analysis/BoxIntegral/Partition/SubboxInduction.lean +++ b/Mathlib/Analysis/BoxIntegral/Partition/SubboxInduction.lean @@ -99,7 +99,7 @@ well-defined. -/ theorem exists_taggedPartition_isHenstock_isSubordinate_homothetic (I : Box ι) (r : (ι → ℝ) → Ioi (0 : ℝ)) : ∃ π : TaggedPrepartition I, π.IsPartition ∧ π.IsHenstock ∧ π.IsSubordinate r ∧ - (∀ J ∈ π, ∃ m : ℕ, ∀ i, (J : _).upper i - J.lower i = (I.upper i - I.lower i) / 2 ^ m) ∧ + (∀ J ∈ π, ∃ m : ℕ, ∀ i, (J :).upper i - J.lower i = (I.upper i - I.lower i) / 2 ^ m) ∧ π.distortion = I.distortion := by refine subbox_induction_on I (fun J _ hJ => ?_) fun z _ => ?_ · choose! πi hP hHen hr Hn _ using hJ @@ -107,7 +107,7 @@ theorem exists_taggedPartition_isHenstock_isSubordinate_homothetic (I : Box ι) have hP : ((splitCenter J).biUnionTagged πi).IsPartition := (isPartition_splitCenter _).biUnionTagged hP have hsub : ∀ J' ∈ (splitCenter J).biUnionTagged πi, ∃ n : ℕ, ∀ i, - (J' : _).upper i - J'.lower i = (J.upper i - J.lower i) / 2 ^ n := by + (J' :).upper i - J'.lower i = (J.upper i - J.lower i) / 2 ^ n := by intro J' hJ' rcases (splitCenter J).mem_biUnionTagged.1 hJ' with ⟨J₁, h₁, h₂⟩ refine ⟨n J₁ J' + 1, fun i => ?_⟩ diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean index ea55c06fddf1e..0cf2fabda7a17 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean @@ -724,7 +724,7 @@ variable [Ring A] [StarRing A] [Algebra R A] [ContinuousFunctionalCalculus R p] lemma isUnit_cfc_iff (f : R → R) (a : A) (hf : ContinuousOn f (spectrum R a) := by cfc_cont_tac) (ha : p a := by cfc_tac) : IsUnit (cfc f a) ↔ ∀ x ∈ spectrum R a, f x ≠ 0 := by rw [← spectrum.zero_not_mem_iff R, cfc_map_spectrum ..] - aesop + simp alias ⟨_, isUnit_cfc⟩ := isUnit_cfc_iff diff --git a/Mathlib/Analysis/CStarAlgebra/Matrix.lean b/Mathlib/Analysis/CStarAlgebra/Matrix.lean index 5db2ac0e4cf48..16c1080e57b78 100644 --- a/Mathlib/Analysis/CStarAlgebra/Matrix.lean +++ b/Mathlib/Analysis/CStarAlgebra/Matrix.lean @@ -171,53 +171,35 @@ scoped[Matrix.L2OpNorm] attribute [instance] Matrix.instL2OpNormedAddCommGroup lemma l2_opNorm_def (A : Matrix m n 𝕜) : ‖A‖ = ‖(toEuclideanLin (𝕜 := 𝕜) (m := m) (n := n)).trans toContinuousLinearMap A‖ := rfl -@[deprecated (since := "2024-02-02")] alias l2_op_norm_def := l2_opNorm_def - lemma l2_opNNNorm_def (A : Matrix m n 𝕜) : ‖A‖₊ = ‖(toEuclideanLin (𝕜 := 𝕜) (m := m) (n := n)).trans toContinuousLinearMap A‖₊ := rfl -@[deprecated (since := "2024-02-02")] alias l2_op_nnnorm_def := l2_opNNNorm_def - lemma l2_opNorm_conjTranspose [DecidableEq m] (A : Matrix m n 𝕜) : ‖Aᴴ‖ = ‖A‖ := by rw [l2_opNorm_def, toEuclideanLin_eq_toLin_orthonormal, LinearEquiv.trans_apply, toLin_conjTranspose, adjoint_toContinuousLinearMap] exact ContinuousLinearMap.adjoint.norm_map _ -@[deprecated (since := "2024-02-02")] alias l2_op_norm_conjTranspose := l2_opNorm_conjTranspose - lemma l2_opNNNorm_conjTranspose [DecidableEq m] (A : Matrix m n 𝕜) : ‖Aᴴ‖₊ = ‖A‖₊ := Subtype.ext <| l2_opNorm_conjTranspose _ -@[deprecated (since := "2024-02-02")] alias l2_op_nnnorm_conjTranspose := l2_opNNNorm_conjTranspose - lemma l2_opNorm_conjTranspose_mul_self (A : Matrix m n 𝕜) : ‖Aᴴ * A‖ = ‖A‖ * ‖A‖ := by classical rw [l2_opNorm_def, toEuclideanLin_eq_toLin_orthonormal, LinearEquiv.trans_apply, Matrix.toLin_mul (v₂ := (EuclideanSpace.basisFun m 𝕜).toBasis), toLin_conjTranspose] exact ContinuousLinearMap.norm_adjoint_comp_self _ -@[deprecated (since := "2024-02-02")] -alias l2_op_norm_conjTranspose_mul_self := l2_opNorm_conjTranspose_mul_self - lemma l2_opNNNorm_conjTranspose_mul_self (A : Matrix m n 𝕜) : ‖Aᴴ * A‖₊ = ‖A‖₊ * ‖A‖₊ := Subtype.ext <| l2_opNorm_conjTranspose_mul_self _ -@[deprecated (since := "2024-02-02")] -alias l2_op_nnnorm_conjTranspose_mul_self := l2_opNNNorm_conjTranspose_mul_self - -- note: with only a type ascription in the left-hand side, Lean picks the wrong norm. lemma l2_opNorm_mulVec (A : Matrix m n 𝕜) (x : EuclideanSpace 𝕜 n) : ‖(EuclideanSpace.equiv m 𝕜).symm <| A *ᵥ x‖ ≤ ‖A‖ * ‖x‖ := toEuclideanLin (n := n) (m := m) (𝕜 := 𝕜) |>.trans toContinuousLinearMap A |>.le_opNorm x -@[deprecated (since := "2024-02-02")] alias l2_op_norm_mulVec := l2_opNorm_mulVec - lemma l2_opNNNorm_mulVec (A : Matrix m n 𝕜) (x : EuclideanSpace 𝕜 n) : ‖(EuclideanSpace.equiv m 𝕜).symm <| A *ᵥ x‖₊ ≤ ‖A‖₊ * ‖x‖₊ := A.l2_opNorm_mulVec x -@[deprecated (since := "2024-02-02")] alias l2_op_nnnorm_mulVec := l2_opNNNorm_mulVec - lemma l2_opNorm_mul (A : Matrix m n 𝕜) (B : Matrix n l 𝕜) : ‖A * B‖ ≤ ‖A‖ * ‖B‖ := by simp only [l2_opNorm_def] @@ -227,13 +209,9 @@ lemma l2_opNorm_mul (A : Matrix m n 𝕜) (B : Matrix n l 𝕜) : ext1 x exact congr($(Matrix.toLin'_mul A B) x) -@[deprecated (since := "2024-02-02")] alias l2_op_norm_mul := l2_opNorm_mul - lemma l2_opNNNorm_mul (A : Matrix m n 𝕜) (B : Matrix n l 𝕜) : ‖A * B‖₊ ≤ ‖A‖₊ * ‖B‖₊ := l2_opNorm_mul A B -@[deprecated (since := "2024-02-02")] alias l2_op_nnnorm_mul := l2_opNNNorm_mul - /-- The normed algebra structure on `Matrix n n 𝕜` arising from the operator norm given by the identification with (continuous) linear endmorphisms of `EuclideanSpace 𝕜 n`. -/ def instL2OpNormedSpace : NormedSpace 𝕜 (Matrix m n 𝕜) where diff --git a/Mathlib/Analysis/CStarAlgebra/Multiplier.lean b/Mathlib/Analysis/CStarAlgebra/Multiplier.lean index 900b2eb10e19f..97b6d4406dc4b 100644 --- a/Mathlib/Analysis/CStarAlgebra/Multiplier.lean +++ b/Mathlib/Analysis/CStarAlgebra/Multiplier.lean @@ -530,7 +530,7 @@ theorem nnnorm_def' (a : 𝓜(𝕜, A)) : ‖a‖₊ = ‖toProdMulOppositeHom a instance instNormedSpace : NormedSpace 𝕜 𝓜(𝕜, A) := { DoubleCentralizer.instModule with - norm_smul_le := fun k a => (norm_smul_le k a.toProdMulOpposite : _) } + norm_smul_le := fun k a => (norm_smul_le k a.toProdMulOpposite :) } instance instNormedAlgebra : NormedAlgebra 𝕜 𝓜(𝕜, A) := { DoubleCentralizer.instAlgebra, DoubleCentralizer.instNormedSpace with } diff --git a/Mathlib/Analysis/CStarAlgebra/Unitization.lean b/Mathlib/Analysis/CStarAlgebra/Unitization.lean index 8771d6ba22cae..caa411ee70756 100644 --- a/Mathlib/Analysis/CStarAlgebra/Unitization.lean +++ b/Mathlib/Analysis/CStarAlgebra/Unitization.lean @@ -38,13 +38,9 @@ lemma opNorm_mul_flip_apply (a : E) : ‖(mul 𝕜 E).flip a‖ = ‖a‖ := by _ ≤ ‖(mul 𝕜 E).flip a‖ * ‖b‖ := by simpa only [flip_apply, mul_apply', norm_star] using le_opNorm ((mul 𝕜 E).flip a) (star b) -@[deprecated (since := "2024-02-02")] alias op_norm_mul_flip_apply := opNorm_mul_flip_apply - lemma opNNNorm_mul_flip_apply (a : E) : ‖(mul 𝕜 E).flip a‖₊ = ‖a‖₊ := Subtype.ext (opNorm_mul_flip_apply 𝕜 a) -@[deprecated (since := "2024-02-02")] alias op_nnnorm_mul_flip_apply := opNNNorm_mul_flip_apply - variable (E) lemma isometry_mul_flip : Isometry (mul 𝕜 E).flip := diff --git a/Mathlib/Analysis/Calculus/BumpFunction/FiniteDimension.lean b/Mathlib/Analysis/Calculus/BumpFunction/FiniteDimension.lean index 70091c5af4b09..3642e77febe4b 100644 --- a/Mathlib/Analysis/Calculus/BumpFunction/FiniteDimension.lean +++ b/Mathlib/Analysis/Calculus/BumpFunction/FiniteDimension.lean @@ -150,10 +150,10 @@ theorem IsOpen.exists_smooth_support_eq {s : Set E} (hs : IsOpen s) : calc ‖iteratedFDeriv ℝ i ((M⁻¹ * δ n) • g n) x‖ = ‖(M⁻¹ * δ n) • iteratedFDeriv ℝ i (g n) x‖ := by rw [iteratedFDeriv_const_smul_apply] - exact (g_smooth n).of_le (mod_cast le_top) + exact (g_smooth n).contDiffAt.of_le (mod_cast le_top) _ = M⁻¹ * δ n * ‖iteratedFDeriv ℝ i (g n) x‖ := by rw [norm_smul _ (iteratedFDeriv ℝ i (g n) x), Real.norm_of_nonneg]; positivity - _ ≤ M⁻¹ * δ n * M := (mul_le_mul_of_nonneg_left ((hR i x).trans (IR i hi)) (by positivity)) + _ ≤ M⁻¹ * δ n * M := by gcongr; exact (hR i x).trans (IR i hi) _ = δ n := by field_simp choose r rpos hr using this have S : ∀ x, Summable fun n => (r n • g n) x := fun x ↦ by diff --git a/Mathlib/Analysis/Calculus/ContDiff/Basic.lean b/Mathlib/Analysis/Calculus/ContDiff/Basic.lean index 87efceee82a5c..75de38f4c2f73 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/Basic.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/Basic.lean @@ -210,19 +210,22 @@ theorem ContDiff.continuousLinearMap_comp {f : E → F} (g : F →L[𝕜] G) (hf /-- The iterated derivative within a set of the composition with a linear map on the left is obtained by applying the linear map to the iterated derivative. -/ theorem ContinuousLinearMap.iteratedFDerivWithin_comp_left {f : E → F} (g : F →L[𝕜] G) - (hf : ContDiffOn 𝕜 n f s) (hs : UniqueDiffOn 𝕜 s) (hx : x ∈ s) {i : ℕ} (hi : i ≤ n) : + (hf : ContDiffWithinAt 𝕜 n f s x) (hs : UniqueDiffOn 𝕜 s) (hx : x ∈ s) {i : ℕ} (hi : i ≤ n) : iteratedFDerivWithin 𝕜 i (g ∘ f) s x = - g.compContinuousMultilinearMap (iteratedFDerivWithin 𝕜 i f s x) := - ((((hf.of_le hi).ftaylorSeriesWithin hs).continuousLinearMap_comp - g).eq_iteratedFDerivWithin_of_uniqueDiffOn le_rfl hs hx).symm + g.compContinuousMultilinearMap (iteratedFDerivWithin 𝕜 i f s x) := by + rcases hf.contDiffOn' hi (by simp) with ⟨U, hU, hxU, hfU⟩ + rw [← iteratedFDerivWithin_inter_open hU hxU, ← iteratedFDerivWithin_inter_open (f := f) hU hxU] + rw [insert_eq_of_mem hx] at hfU + exact .symm <| (hfU.ftaylorSeriesWithin (hs.inter hU)).continuousLinearMap_comp g + |>.eq_iteratedFDerivWithin_of_uniqueDiffOn le_rfl (hs.inter hU) ⟨hx, hxU⟩ /-- The iterated derivative of the composition with a linear map on the left is obtained by applying the linear map to the iterated derivative. -/ theorem ContinuousLinearMap.iteratedFDeriv_comp_left {f : E → F} (g : F →L[𝕜] G) - (hf : ContDiff 𝕜 n f) (x : E) {i : ℕ} (hi : i ≤ n) : + (hf : ContDiffAt 𝕜 n f x) {i : ℕ} (hi : i ≤ n) : iteratedFDeriv 𝕜 i (g ∘ f) x = g.compContinuousMultilinearMap (iteratedFDeriv 𝕜 i f x) := by simp only [← iteratedFDerivWithin_univ] - exact g.iteratedFDerivWithin_comp_left hf.contDiffOn uniqueDiffOn_univ (mem_univ x) hi + exact g.iteratedFDerivWithin_comp_left hf.contDiffWithinAt uniqueDiffOn_univ (mem_univ x) hi /-- The iterated derivative within a set of the composition with a linear equiv on the left is obtained by applying the linear equiv to the iterated derivative. This is true without @@ -251,7 +254,7 @@ theorem ContinuousLinearEquiv.iteratedFDerivWithin_comp_left (g : F ≃L[𝕜] G /-- Composition with a linear isometry on the left preserves the norm of the iterated derivative within a set. -/ theorem LinearIsometry.norm_iteratedFDerivWithin_comp_left {f : E → F} (g : F →ₗᵢ[𝕜] G) - (hf : ContDiffOn 𝕜 n f s) (hs : UniqueDiffOn 𝕜 s) (hx : x ∈ s) {i : ℕ} (hi : i ≤ n) : + (hf : ContDiffWithinAt 𝕜 n f s x) (hs : UniqueDiffOn 𝕜 s) (hx : x ∈ s) {i : ℕ} (hi : i ≤ n) : ‖iteratedFDerivWithin 𝕜 i (g ∘ f) s x‖ = ‖iteratedFDerivWithin 𝕜 i f s x‖ := by have : iteratedFDerivWithin 𝕜 i (g ∘ f) s x = @@ -263,10 +266,10 @@ theorem LinearIsometry.norm_iteratedFDerivWithin_comp_left {f : E → F} (g : F /-- Composition with a linear isometry on the left preserves the norm of the iterated derivative. -/ theorem LinearIsometry.norm_iteratedFDeriv_comp_left {f : E → F} (g : F →ₗᵢ[𝕜] G) - (hf : ContDiff 𝕜 n f) (x : E) {i : ℕ} (hi : i ≤ n) : + (hf : ContDiffAt 𝕜 n f x) {i : ℕ} (hi : i ≤ n) : ‖iteratedFDeriv 𝕜 i (g ∘ f) x‖ = ‖iteratedFDeriv 𝕜 i f x‖ := by simp only [← iteratedFDerivWithin_univ] - exact g.norm_iteratedFDerivWithin_comp_left hf.contDiffOn uniqueDiffOn_univ (mem_univ x) hi + exact g.norm_iteratedFDerivWithin_comp_left hf.contDiffWithinAt uniqueDiffOn_univ (mem_univ x) hi /-- Composition with a linear isometry equiv on the left preserves the norm of the iterated derivative within a set. -/ @@ -700,6 +703,45 @@ theorem ContDiff.comp_contDiffAt {g : F → G} {f : E → F} (x : E) (hg : ContD (hf : ContDiffAt 𝕜 n f x) : ContDiffAt 𝕜 n (g ∘ f) x := hg.comp_contDiffWithinAt hf +theorem iteratedFDerivWithin_comp_of_eventually_mem {t : Set F} + (hg : ContDiffWithinAt 𝕜 n g t (f x)) (hf : ContDiffWithinAt 𝕜 n f s x) + (ht : UniqueDiffOn 𝕜 t) (hs : UniqueDiffOn 𝕜 s) (hxs : x ∈ s) (hst : ∀ᶠ y in 𝓝[s] x, f y ∈ t) + {i : ℕ} (hi : i ≤ n) : + iteratedFDerivWithin 𝕜 i (g ∘ f) s x = + (ftaylorSeriesWithin 𝕜 g t (f x)).taylorComp (ftaylorSeriesWithin 𝕜 f s x) i := by + obtain ⟨u, hxu, huo, hfu, hgu⟩ : ∃ u, x ∈ u ∧ IsOpen u ∧ + HasFTaylorSeriesUpToOn i f (ftaylorSeriesWithin 𝕜 f s) (s ∩ u) ∧ + HasFTaylorSeriesUpToOn i g (ftaylorSeriesWithin 𝕜 g t) (f '' (s ∩ u)) := by + have hxt : f x ∈ t := hst.self_of_nhdsWithin hxs + have hf_tendsto : Tendsto f (𝓝[s] x) (𝓝[t] (f x)) := + tendsto_nhdsWithin_iff.mpr ⟨hf.continuousWithinAt, hst⟩ + have H₁ : ∀ᶠ u in (𝓝[s] x).smallSets, + HasFTaylorSeriesUpToOn i f (ftaylorSeriesWithin 𝕜 f s) u := + hf.eventually_hasFTaylorSeriesUpToOn hs hxs hi + have H₂ : ∀ᶠ u in (𝓝[s] x).smallSets, + HasFTaylorSeriesUpToOn i g (ftaylorSeriesWithin 𝕜 g t) (f '' u) := + hf_tendsto.image_smallSets.eventually (hg.eventually_hasFTaylorSeriesUpToOn ht hxt hi) + rcases (nhdsWithin_basis_open _ _).smallSets.eventually_iff.mp (H₁.and H₂) + with ⟨u, ⟨hxu, huo⟩, hu⟩ + exact ⟨u, hxu, huo, hu (by simp [inter_comm])⟩ + exact .symm <| (hgu.comp hfu (mapsTo_image _ _)).eq_iteratedFDerivWithin_of_uniqueDiffOn le_rfl + (hs.inter huo) ⟨hxs, hxu⟩ |>.trans <| iteratedFDerivWithin_inter_open huo hxu + +theorem iteratedFDerivWithin_comp {t : Set F} (hg : ContDiffWithinAt 𝕜 n g t (f x)) + (hf : ContDiffWithinAt 𝕜 n f s x) (ht : UniqueDiffOn 𝕜 t) (hs : UniqueDiffOn 𝕜 s) + (hx : x ∈ s) (hst : MapsTo f s t) {i : ℕ} (hi : i ≤ n) : + iteratedFDerivWithin 𝕜 i (g ∘ f) s x = + (ftaylorSeriesWithin 𝕜 g t (f x)).taylorComp (ftaylorSeriesWithin 𝕜 f s x) i := + iteratedFDerivWithin_comp_of_eventually_mem hg hf ht hs hx (eventually_mem_nhdsWithin.mono hst) hi + +theorem iteratedFDeriv_comp (hg : ContDiffAt 𝕜 n g (f x)) (hf : ContDiffAt 𝕜 n f x) + {i : ℕ} (hi : i ≤ n) : + iteratedFDeriv 𝕜 i (g ∘ f) x = + (ftaylorSeries 𝕜 g (f x)).taylorComp (ftaylorSeries 𝕜 f x) i := by + simp only [← iteratedFDerivWithin_univ, ← ftaylorSeriesWithin_univ] + exact iteratedFDerivWithin_comp hg.contDiffWithinAt hf.contDiffWithinAt + uniqueDiffOn_univ uniqueDiffOn_univ (mem_univ _) (mapsTo_univ _ _) hi + /-! ### Smoothness of projections -/ @@ -1612,16 +1654,16 @@ theorem ContDiffOn.const_smul {s : Set E} {f : E → F} (c : R) (hf : ContDiffOn variable {i : ℕ} {a : R} -theorem iteratedFDerivWithin_const_smul_apply (hf : ContDiffOn 𝕜 i f s) (hu : UniqueDiffOn 𝕜 s) - (hx : x ∈ s) : iteratedFDerivWithin 𝕜 i (a • f) s x = a • iteratedFDerivWithin 𝕜 i f s x := +theorem iteratedFDerivWithin_const_smul_apply (hf : ContDiffWithinAt 𝕜 i f s x) + (hu : UniqueDiffOn 𝕜 s) (hx : x ∈ s) : + iteratedFDerivWithin 𝕜 i (a • f) s x = a • iteratedFDerivWithin 𝕜 i f s x := (a • (1 : F →L[𝕜] F)).iteratedFDerivWithin_comp_left hf hu hx le_rfl -theorem iteratedFDeriv_const_smul_apply {x : E} (hf : ContDiff 𝕜 i f) : - iteratedFDeriv 𝕜 i (a • f) x = a • iteratedFDeriv 𝕜 i f x := by - simp_rw [← contDiffOn_univ, ← iteratedFDerivWithin_univ] at * - exact iteratedFDerivWithin_const_smul_apply hf uniqueDiffOn_univ (Set.mem_univ _) +theorem iteratedFDeriv_const_smul_apply (hf : ContDiffAt 𝕜 i f x) : + iteratedFDeriv 𝕜 i (a • f) x = a • iteratedFDeriv 𝕜 i f x := + (a • (1 : F →L[𝕜] F)).iteratedFDeriv_comp_left hf le_rfl -theorem iteratedFDeriv_const_smul_apply' {x : E} (hf : ContDiff 𝕜 i f) : +theorem iteratedFDeriv_const_smul_apply' (hf : ContDiffAt 𝕜 i f x) : iteratedFDeriv 𝕜 i (fun x ↦ a • f x) x = a • iteratedFDeriv 𝕜 i f x := iteratedFDeriv_const_smul_apply hf diff --git a/Mathlib/Analysis/Calculus/ContDiff/Bounds.lean b/Mathlib/Analysis/Calculus/ContDiff/Bounds.lean index c869df9c43a51..e5879e948ab45 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/Bounds.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/Bounds.lean @@ -546,7 +546,7 @@ theorem norm_iteratedFDeriv_clm_apply {f : E → F →L[𝕜] G} {g : E → F} { (Set.mem_univ x) hn theorem norm_iteratedFDerivWithin_clm_apply_const {f : E → F →L[𝕜] G} {c : F} {s : Set E} {x : E} - {N : WithTop ℕ∞} {n : ℕ} (hf : ContDiffOn 𝕜 N f s) (hs : UniqueDiffOn 𝕜 s) + {N : WithTop ℕ∞} {n : ℕ} (hf : ContDiffWithinAt 𝕜 N f s x) (hs : UniqueDiffOn 𝕜 s) (hx : x ∈ s) (hn : n ≤ N) : ‖iteratedFDerivWithin 𝕜 n (fun y : E => (f y) c) s x‖ ≤ ‖c‖ * ‖iteratedFDerivWithin 𝕜 n f s x‖ := by @@ -559,10 +559,10 @@ theorem norm_iteratedFDerivWithin_clm_apply_const {f : E → F →L[𝕜] G} {c exact f.le_opNorm c theorem norm_iteratedFDeriv_clm_apply_const {f : E → F →L[𝕜] G} {c : F} {x : E} - {N : WithTop ℕ∞} {n : ℕ} (hf : ContDiff 𝕜 N f) (hn : n ≤ N) : + {N : WithTop ℕ∞} {n : ℕ} (hf : ContDiffAt 𝕜 N f x) (hn : n ≤ N) : ‖iteratedFDeriv 𝕜 n (fun y : E => (f y) c) x‖ ≤ ‖c‖ * ‖iteratedFDeriv 𝕜 n f x‖ := by simp only [← iteratedFDerivWithin_univ] - exact norm_iteratedFDerivWithin_clm_apply_const hf.contDiffOn uniqueDiffOn_univ + exact norm_iteratedFDerivWithin_clm_apply_const hf.contDiffWithinAt uniqueDiffOn_univ (Set.mem_univ x) hn end Apply diff --git a/Mathlib/Analysis/Calculus/ContDiff/Defs.lean b/Mathlib/Analysis/Calculus/ContDiff/Defs.lean index eabcdff56fab9..1224f6b8ffeae 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/Defs.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/Defs.lean @@ -242,7 +242,7 @@ alias Filter.EventuallyEq.contDiffWithinAt_iff := Filter.EventuallyEq.congr_cont theorem ContDiffWithinAt.congr_of_eventuallyEq_insert (h : ContDiffWithinAt 𝕜 n f s x) (h₁ : f₁ =ᶠ[𝓝[insert x s] x] f) : ContDiffWithinAt 𝕜 n f₁ s x := h.congr_of_eventuallyEq (nhdsWithin_mono x (subset_insert x s) h₁) - (mem_of_mem_nhdsWithin (mem_insert x s) h₁ : _) + (mem_of_mem_nhdsWithin (mem_insert x s) h₁ :) theorem Filter.EventuallyEq.congr_contDiffWithinAt_of_insert (h₁ : f₁ =ᶠ[𝓝[insert x s] x] f) : ContDiffWithinAt 𝕜 n f₁ s x ↔ ContDiffWithinAt 𝕜 n f s x := @@ -693,6 +693,18 @@ theorem iteratedFDerivWithin_subset {n : ℕ} (st : s ⊆ t) (hs : UniqueDiffOn iteratedFDerivWithin 𝕜 n f s x = iteratedFDerivWithin 𝕜 n f t x := (((h.ftaylorSeriesWithin ht).mono st).eq_iteratedFDerivWithin_of_uniqueDiffOn le_rfl hs hx).symm +theorem ContDiffWithinAt.eventually_hasFTaylorSeriesUpToOn {f : E → F} {s : Set E} {a : E} + (h : ContDiffWithinAt 𝕜 n f s a) (hs : UniqueDiffOn 𝕜 s) (ha : a ∈ s) {m : ℕ} (hm : m ≤ n) : + ∀ᶠ t in (𝓝[s] a).smallSets, HasFTaylorSeriesUpToOn m f (ftaylorSeriesWithin 𝕜 f s) t := by + rcases h.contDiffOn' hm (by simp) with ⟨U, hUo, haU, hfU⟩ + have : ∀ᶠ t in (𝓝[s] a).smallSets, t ⊆ s ∩ U := by + rw [eventually_smallSets_subset] + exact inter_mem_nhdsWithin _ <| hUo.mem_nhds haU + refine this.mono fun t ht ↦ .mono ?_ ht + rw [insert_eq_of_mem ha] at hfU + refine (hfU.ftaylorSeriesWithin (hs.inter hUo)).congr_series fun k hk x hx ↦ ?_ + exact iteratedFDerivWithin_inter_open hUo hx.2 + /-- On a set with unique differentiability, an analytic function is automatically `C^ω`, as its successive derivatives are also analytic. This does not require completeness of the space. See also `AnalyticOn.contDiffOn_of_completeSpace`.-/ diff --git a/Mathlib/Analysis/Calculus/ContDiff/FTaylorSeries.lean b/Mathlib/Analysis/Calculus/ContDiff/FTaylorSeries.lean index 400df9cb52915..7cc8a905e12ec 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/FTaylorSeries.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/FTaylorSeries.lean @@ -412,9 +412,14 @@ theorem iteratedFDerivWithin_zero_eq_comp : iteratedFDerivWithin 𝕜 0 f s = (continuousMultilinearCurryFin0 𝕜 E F).symm ∘ f := rfl +@[simp] +theorem dist_iteratedFDerivWithin_zero (f : E → F) (s : Set E) (x : E) + (g : E → F) (t : Set E) (y : E) : + dist (iteratedFDerivWithin 𝕜 0 f s x) (iteratedFDerivWithin 𝕜 0 g t y) = dist (f x) (g y) := by + simp only [iteratedFDerivWithin_zero_eq_comp, comp_apply, LinearIsometryEquiv.dist_map] + @[simp] theorem norm_iteratedFDerivWithin_zero : ‖iteratedFDerivWithin 𝕜 0 f s x‖ = ‖f x‖ := by - -- Porting note: added `comp_apply`. rw [iteratedFDerivWithin_zero_eq_comp, comp_apply, LinearIsometryEquiv.norm_map] theorem iteratedFDerivWithin_succ_apply_left {n : ℕ} (m : Fin (n + 1) → E) : @@ -442,6 +447,23 @@ theorem norm_fderivWithin_iteratedFDerivWithin {n : ℕ} : -- Porting note: added `comp_apply`. rw [iteratedFDerivWithin_succ_eq_comp_left, comp_apply, LinearIsometryEquiv.norm_map] +@[simp] +theorem dist_iteratedFDerivWithin_one (f g : E → F) {y} + (hsx : UniqueDiffWithinAt 𝕜 s x) (hyt : UniqueDiffWithinAt 𝕜 t y) : + dist (iteratedFDerivWithin 𝕜 1 f s x) (iteratedFDerivWithin 𝕜 1 g t y) + = dist (fderivWithin 𝕜 f s x) (fderivWithin 𝕜 g t y) := by + simp only [iteratedFDerivWithin_succ_eq_comp_left, comp_apply, + LinearIsometryEquiv.dist_map, iteratedFDerivWithin_zero_eq_comp, + LinearIsometryEquiv.comp_fderivWithin, hsx, hyt] + apply (continuousMultilinearCurryFin0 𝕜 E F).symm.toLinearIsometry.postcomp.dist_map + +@[simp] +theorem norm_iteratedFDerivWithin_one (f : E → F) (h : UniqueDiffWithinAt 𝕜 s x) : + ‖iteratedFDerivWithin 𝕜 1 f s x‖ = ‖fderivWithin 𝕜 f s x‖ := by + simp only [← norm_fderivWithin_iteratedFDerivWithin, + iteratedFDerivWithin_zero_eq_comp, LinearIsometryEquiv.comp_fderivWithin _ h] + apply (continuousMultilinearCurryFin0 𝕜 E F).symm.toLinearIsometry.norm_toContinuousLinearMap_comp + theorem iteratedFDerivWithin_succ_apply_right {n : ℕ} (hs : UniqueDiffOn 𝕜 s) (hx : x ∈ s) (m : Fin (n + 1) → E) : (iteratedFDerivWithin 𝕜 (n + 1) f s x : (Fin (n + 1) → E) → F) m = @@ -630,10 +652,6 @@ theorem HasFTaylorSeriesUpToOn.eq_iteratedFDerivWithin_of_uniqueDiffOn rw [iteratedFDerivWithin_succ_eq_comp_left, Function.comp_apply, this.fderivWithin (hs x hx)] exact (ContinuousMultilinearMap.uncurry_curryLeft _).symm -@[deprecated (since := "2024-03-28")] -alias HasFTaylorSeriesUpToOn.eq_ftaylor_series_of_uniqueDiffOn := - HasFTaylorSeriesUpToOn.eq_iteratedFDerivWithin_of_uniqueDiffOn - /-- The iterated derivative commutes with shifting the function by a constant on the left. -/ lemma iteratedFDerivWithin_comp_add_left' (n : ℕ) (a : E) : iteratedFDerivWithin 𝕜 n (fun z ↦ f (a + z)) s = diff --git a/Mathlib/Analysis/Calculus/Deriv/Basic.lean b/Mathlib/Analysis/Calculus/Deriv/Basic.lean index d0e6db0c4bcf9..64f66f524e715 100644 --- a/Mathlib/Analysis/Calculus/Deriv/Basic.lean +++ b/Mathlib/Analysis/Calculus/Deriv/Basic.lean @@ -562,7 +562,7 @@ theorem HasDerivWithinAt.congr_deriv (h : HasDerivWithinAt f f' s x) (h' : f' = theorem HasDerivAt.congr_of_eventuallyEq (h : HasDerivAt f f' x) (h₁ : f₁ =ᶠ[𝓝 x] f) : HasDerivAt f₁ f' x := - HasDerivAtFilter.congr_of_eventuallyEq h h₁ (mem_of_mem_nhds h₁ : _) + HasDerivAtFilter.congr_of_eventuallyEq h h₁ (mem_of_mem_nhds h₁ :) theorem Filter.EventuallyEq.hasDerivAt_iff (h : f₀ =ᶠ[𝓝 x] f₁) : HasDerivAt f₀ f' x ↔ HasDerivAt f₁ f' x := diff --git a/Mathlib/Analysis/Calculus/FDeriv/Basic.lean b/Mathlib/Analysis/Calculus/FDeriv/Basic.lean index c22d0f09c8255..e49d2ade9920f 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Basic.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Basic.lean @@ -854,7 +854,7 @@ theorem HasFDerivWithinAt.congr_of_eventuallyEq (h : HasFDerivWithinAt f f' s x) theorem HasFDerivAt.congr_of_eventuallyEq (h : HasFDerivAt f f' x) (h₁ : f₁ =ᶠ[𝓝 x] f) : HasFDerivAt f₁ f' x := - HasFDerivAtFilter.congr_of_eventuallyEq h h₁ (mem_of_mem_nhds h₁ : _) + HasFDerivAtFilter.congr_of_eventuallyEq h h₁ (mem_of_mem_nhds h₁ :) theorem DifferentiableWithinAt.congr_mono (h : DifferentiableWithinAt 𝕜 f s x) (ht : EqOn f₁ f t) (hx : f₁ x = f x) (h₁ : t ⊆ s) : DifferentiableWithinAt 𝕜 f₁ t x := diff --git a/Mathlib/Analysis/Calculus/FDeriv/Equiv.lean b/Mathlib/Analysis/Calculus/FDeriv/Equiv.lean index b859dc97ba096..6445fa96ecce5 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Equiv.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Equiv.lean @@ -322,6 +322,30 @@ theorem comp_fderiv' {f : G → E} : end LinearIsometryEquiv +/-- If `f (g y) = y` for `y` in a neighborhood of `a` within `t`, +`g` maps a neighborhood of `a` within `t` to a neighborhood of `g a` within `s`, +and `f` has an invertible derivative `f'` at `g a` within `s`, +then `g` has the derivative `f'⁻¹` at `a` within `t`. + +This is one of the easy parts of the inverse function theorem: it assumes that we already have an +inverse function. -/ +theorem HasFDerivWithinAt.of_local_left_inverse {g : F → E} {f' : E ≃L[𝕜] F} {a : F} {t : Set F} + (hg : Tendsto g (𝓝[t] a) (𝓝[s] (g a))) (hf : HasFDerivWithinAt f (f' : E →L[𝕜] F) s (g a)) + (ha : a ∈ t) (hfg : ∀ᶠ y in 𝓝[t] a, f (g y) = y) : + HasFDerivWithinAt g (f'.symm : F →L[𝕜] E) t a := by + have : (fun x : F => g x - g a - f'.symm (x - a)) =O[𝓝[t] a] + fun x : F => f' (g x - g a) - (x - a) := + ((f'.symm : F →L[𝕜] E).isBigO_comp _ _).congr (fun x ↦ by simp) fun _ ↦ rfl + refine .of_isLittleO <| this.trans_isLittleO ?_ + clear this + refine ((hf.isLittleO.comp_tendsto hg).symm.congr' (hfg.mono ?_) .rfl).trans_isBigO ?_ + · intro p hp + simp [hp, hfg.self_of_nhdsWithin ha] + · refine ((hf.isBigO_sub_rev f'.antilipschitz).comp_tendsto hg).congr' + (Eventually.of_forall fun _ => rfl) (hfg.mono ?_) + rintro p hp + simp only [(· ∘ ·), hp, hfg.self_of_nhdsWithin ha] + /-- If `f (g y) = y` for `y` in some neighborhood of `a`, `g` is continuous at `a`, and `f` has an invertible derivative `f'` at `g a` in the strict sense, then `g` has the derivative `f'⁻¹` at `a` in the strict sense. @@ -357,19 +381,8 @@ an inverse function. -/ theorem HasFDerivAt.of_local_left_inverse {f : E → F} {f' : E ≃L[𝕜] F} {g : F → E} {a : F} (hg : ContinuousAt g a) (hf : HasFDerivAt f (f' : E →L[𝕜] F) (g a)) (hfg : ∀ᶠ y in 𝓝 a, f (g y) = y) : HasFDerivAt g (f'.symm : F →L[𝕜] E) a := by - have : (fun x : F => g x - g a - f'.symm (x - a)) =O[𝓝 a] - fun x : F => f' (g x - g a) - (x - a) := by - refine ((f'.symm : F →L[𝕜] E).isBigO_comp _ _).congr (fun x => ?_) fun _ => rfl - simp - refine HasFDerivAtFilter.of_isLittleO <| this.trans_isLittleO ?_ - clear this - refine ((hf.isLittleO.comp_tendsto hg).symm.congr' (hfg.mono ?_) .rfl).trans_isBigO ?_ - · intro p hp - simp [hp, hfg.self_of_nhds] - · refine ((hf.isBigO_sub_rev f'.antilipschitz).comp_tendsto hg).congr' - (Eventually.of_forall fun _ => rfl) (hfg.mono ?_) - rintro p hp - simp only [(· ∘ ·), hp, hfg.self_of_nhds] + simp only [← hasFDerivWithinAt_univ, ← nhdsWithin_univ] at hf hfg ⊢ + exact hf.of_local_left_inverse (.inf hg (by simp)) (mem_univ _) hfg /-- If `f` is a partial homeomorphism defined on a neighbourhood of `f.symm a`, and `f` has an invertible derivative `f'` in the sense of strict differentiability at `f.symm a`, then `f.symm` has diff --git a/Mathlib/Analysis/Calculus/FormalMultilinearSeries.lean b/Mathlib/Analysis/Calculus/FormalMultilinearSeries.lean index 4b12e24cee9e5..4254d8378f1d6 100644 --- a/Mathlib/Analysis/Calculus/FormalMultilinearSeries.lean +++ b/Mathlib/Analysis/Calculus/FormalMultilinearSeries.lean @@ -81,6 +81,10 @@ theorem add_apply (p q : FormalMultilinearSeries 𝕜 E F) (n : ℕ) : (p + q) n @[simp] theorem sub_apply (p q : FormalMultilinearSeries 𝕜 E F) (n : ℕ) : (p - q) n = p n - q n := rfl +@[simp] +theorem smul_apply [Semiring 𝕜'] [Module 𝕜' F] [ContinuousConstSMul 𝕜' F] [SMulCommClass 𝕜 𝕜' F] + (f : FormalMultilinearSeries 𝕜 E F) (n : ℕ) (a : 𝕜') : (a • f) n = a • f n := rfl + @[ext] protected theorem ext {p q : FormalMultilinearSeries 𝕜 E F} (h : ∀ n, p n = q n) : p = q := funext h diff --git a/Mathlib/Analysis/Calculus/Gradient/Basic.lean b/Mathlib/Analysis/Calculus/Gradient/Basic.lean index 5ebf3d941d6a8..5c6489419a6d0 100644 --- a/Mathlib/Analysis/Calculus/Gradient/Basic.lean +++ b/Mathlib/Analysis/Calculus/Gradient/Basic.lean @@ -284,7 +284,7 @@ theorem HasGradientWithinAt.congr_of_eventuallyEq_of_mem (h : HasGradientWithinA theorem HasGradientAt.congr_of_eventuallyEq (h : HasGradientAt f f' x) (h₁ : f₁ =ᶠ[𝓝 x] f) : HasGradientAt f₁ f' x := - HasGradientAtFilter.congr_of_eventuallyEq h h₁ (mem_of_mem_nhds h₁ : _) + HasGradientAtFilter.congr_of_eventuallyEq h h₁ (mem_of_mem_nhds h₁ :) theorem Filter.EventuallyEq.gradient_eq (hL : f₁ =ᶠ[𝓝 x] f) : ∇ f₁ x = ∇ f x := by unfold gradient diff --git a/Mathlib/Analysis/Calculus/InverseFunctionTheorem/Deriv.lean b/Mathlib/Analysis/Calculus/InverseFunctionTheorem/Deriv.lean index 48cbff9518b9a..be38e1e55e5cc 100644 --- a/Mathlib/Analysis/Calculus/InverseFunctionTheorem/Deriv.lean +++ b/Mathlib/Analysis/Calculus/InverseFunctionTheorem/Deriv.lean @@ -49,5 +49,3 @@ variable {f} theorem isOpenMap_of_hasStrictDerivAt {f' : 𝕜 → 𝕜} (hf : ∀ x, HasStrictDerivAt f (f' x) x) (h0 : ∀ x, f' x ≠ 0) : IsOpenMap f := isOpenMap_iff_nhds_le.2 fun x => ((hf x).map_nhds_eq (h0 x)).ge -@[deprecated (since := "2024-03-23")] -alias open_map_of_strict_deriv := isOpenMap_of_hasStrictDerivAt diff --git a/Mathlib/Analysis/Calculus/InverseFunctionTheorem/FDeriv.lean b/Mathlib/Analysis/Calculus/InverseFunctionTheorem/FDeriv.lean index 06f438efc45a0..7f5ef4e24ef29 100644 --- a/Mathlib/Analysis/Calculus/InverseFunctionTheorem/FDeriv.lean +++ b/Mathlib/Analysis/Calculus/InverseFunctionTheorem/FDeriv.lean @@ -194,5 +194,3 @@ end HasStrictFDerivAt theorem isOpenMap_of_hasStrictFDerivAt_equiv [CompleteSpace E] {f : E → F} {f' : E → E ≃L[𝕜] F} (hf : ∀ x, HasStrictFDerivAt f (f' x : E →L[𝕜] F) x) : IsOpenMap f := isOpenMap_iff_nhds_le.2 fun x => (hf x).map_nhds_eq_of_equiv.ge -@[deprecated (since := "2024-03-23")] -alias open_map_of_strict_fderiv_equiv := isOpenMap_of_hasStrictFDerivAt_equiv diff --git a/Mathlib/Analysis/Calculus/IteratedDeriv/Lemmas.lean b/Mathlib/Analysis/Calculus/IteratedDeriv/Lemmas.lean index c10db8c308277..01f21bd7293eb 100644 --- a/Mathlib/Analysis/Calculus/IteratedDeriv/Lemmas.lean +++ b/Mathlib/Analysis/Calculus/IteratedDeriv/Lemmas.lean @@ -65,13 +65,13 @@ theorem iteratedDerivWithin_const_sub (hn : 0 < n) (c : F) : @[deprecated (since := "2024-12-10")] alias iteratedDerivWithin_const_neg := iteratedDerivWithin_const_sub -theorem iteratedDerivWithin_const_smul (c : R) (hf : ContDiffOn 𝕜 n f s) : +theorem iteratedDerivWithin_const_smul (c : R) (hf : ContDiffWithinAt 𝕜 n f s x) : iteratedDerivWithin n (c • f) s x = c • iteratedDerivWithin n f s x := by simp_rw [iteratedDerivWithin] rw [iteratedFDerivWithin_const_smul_apply hf h hx] simp only [ContinuousMultilinearMap.smul_apply] -theorem iteratedDerivWithin_const_mul (c : 𝕜) {f : 𝕜 → 𝕜} (hf : ContDiffOn 𝕜 n f s) : +theorem iteratedDerivWithin_const_mul (c : 𝕜) {f : 𝕜 → 𝕜} (hf : ContDiffWithinAt 𝕜 n f s x) : iteratedDerivWithin n (fun z => c * f z) s x = c * iteratedDerivWithin n f s x := by simpa using iteratedDerivWithin_const_smul (F := 𝕜) hx h c hf diff --git a/Mathlib/Analysis/Calculus/MeanValue.lean b/Mathlib/Analysis/Calculus/MeanValue.lean index 8194a4e3aa8f5..ee6d474ded239 100644 --- a/Mathlib/Analysis/Calculus/MeanValue.lean +++ b/Mathlib/Analysis/Calculus/MeanValue.lean @@ -1008,9 +1008,6 @@ lemma strictMonoOn_of_hasDerivWithinAt_pos {D : Set ℝ} (hD : Convex ℝ D) {f strictMonoOn_of_deriv_pos hD hf fun x hx ↦ by rw [deriv_eqOn isOpen_interior hf' hx]; exact hf'₀ _ hx -@[deprecated (since := "2024-03-02")] -alias StrictMonoOn_of_hasDerivWithinAt_pos := strictMonoOn_of_hasDerivWithinAt_pos - /-- Let `f : ℝ → ℝ` be a differentiable function. If `f'` is strictly positive, then `f` is a strictly monotone function. -/ lemma strictMono_of_hasDerivAt_pos {f f' : ℝ → ℝ} (hf : ∀ x, HasDerivAt f (f' x) x) @@ -1080,18 +1077,12 @@ lemma strictAntiOn_of_hasDerivWithinAt_neg {D : Set ℝ} (hD : Convex ℝ D) {f strictAntiOn_of_deriv_neg hD hf fun x hx ↦ by rw [deriv_eqOn isOpen_interior hf' hx]; exact hf'₀ _ hx -@[deprecated (since := "2024-03-02")] -alias StrictAntiOn_of_hasDerivWithinAt_pos := strictAntiOn_of_hasDerivWithinAt_neg - /-- Let `f : ℝ → ℝ` be a differentiable function. If `f'` is strictly positive, then `f` is a strictly monotone function. -/ lemma strictAnti_of_hasDerivAt_neg {f f' : ℝ → ℝ} (hf : ∀ x, HasDerivAt f (f' x) x) (hf' : ∀ x, f' x < 0) : StrictAnti f := strictAnti_of_deriv_neg fun x ↦ by rw [(hf _).deriv]; exact hf' _ -@[deprecated (since := "2024-03-02")] -alias strictAnti_of_hasDerivAt_pos := strictAnti_of_hasDerivAt_neg - /-- Let `f` be a function continuous on a convex (or, equivalently, connected) subset `D` of the real line. If `f` is differentiable on the interior of `D` and `f'` is nonpositive, then `f` is an antitone function on `D`. -/ diff --git a/Mathlib/Analysis/Complex/Basic.lean b/Mathlib/Analysis/Complex/Basic.lean index 05dc52915295f..73c4f7ae5cdfa 100644 --- a/Mathlib/Analysis/Complex/Basic.lean +++ b/Mathlib/Analysis/Complex/Basic.lean @@ -154,10 +154,10 @@ theorem comap_abs_nhds_zero : comap abs (𝓝 0) = 𝓝 0 := @[simp 1100, norm_cast] lemma nnnorm_ratCast (q : ℚ) : ‖(q : ℂ)‖₊ = ‖(q : ℝ)‖₊ := nnnorm_real q @[simp 1100] lemma norm_ofNat (n : ℕ) [n.AtLeastTwo] : - ‖(no_index (OfNat.ofNat n) : ℂ)‖ = OfNat.ofNat n := norm_natCast n + ‖(ofNat(n) : ℂ)‖ = OfNat.ofNat n := norm_natCast n @[simp 1100] lemma nnnorm_ofNat (n : ℕ) [n.AtLeastTwo] : - ‖(no_index (OfNat.ofNat n) : ℂ)‖₊ = OfNat.ofNat n := nnnorm_natCast n + ‖(ofNat(n) : ℂ)‖₊ = OfNat.ofNat n := nnnorm_natCast n @[deprecated (since := "2024-08-25")] alias norm_nat := norm_natCast @[deprecated (since := "2024-08-25")] alias norm_int := norm_intCast @@ -693,7 +693,7 @@ lemma natCast_mem_slitPlane {n : ℕ} : ↑n ∈ slitPlane ↔ n ≠ 0 := by alias nat_cast_mem_slitPlane := natCast_mem_slitPlane @[simp] -lemma ofNat_mem_slitPlane (n : ℕ) [n.AtLeastTwo] : no_index (OfNat.ofNat n) ∈ slitPlane := +lemma ofNat_mem_slitPlane (n : ℕ) [n.AtLeastTwo] : ofNat(n) ∈ slitPlane := natCast_mem_slitPlane.2 (NeZero.ne n) lemma mem_slitPlane_iff_not_le_zero {z : ℂ} : z ∈ slitPlane ↔ ¬z ≤ 0 := diff --git a/Mathlib/Analysis/Complex/Polynomial/Basic.lean b/Mathlib/Analysis/Complex/Polynomial/Basic.lean index 82e38a8adc578..706b89c25be7f 100644 --- a/Mathlib/Analysis/Complex/Polynomial/Basic.lean +++ b/Mathlib/Analysis/Complex/Polynomial/Basic.lean @@ -216,6 +216,3 @@ lemma Irreducible.degree_le_two {p : ℝ[X]} (hp : Irreducible p) : degree p ≤ /-- An irreducible real polynomial has natural degree at most two. -/ lemma Irreducible.natDegree_le_two {p : ℝ[X]} (hp : Irreducible p) : natDegree p ≤ 2 := natDegree_le_iff_degree_le.2 hp.degree_le_two - -@[deprecated (since := "2024-02-18")] -alias Irreducible.nat_degree_le_two := Irreducible.natDegree_le_two diff --git a/Mathlib/Analysis/Convolution.lean b/Mathlib/Analysis/Convolution.lean index 9363b3cee07de..9834adcc2007f 100644 --- a/Mathlib/Analysis/Convolution.lean +++ b/Mathlib/Analysis/Convolution.lean @@ -931,7 +931,7 @@ variable [NormedAddCommGroup G] [BorelSpace G] theorem convolution_precompR_apply {g : G → E'' →L[𝕜] E'} (hf : LocallyIntegrable f μ) (hcg : HasCompactSupport g) (hg : Continuous g) (x₀ : G) (x : E'') : (f ⋆[L.precompR E'', μ] g) x₀ x = (f ⋆[L, μ] fun a => g a x) x₀ := by - have := hcg.convolutionExists_right (L.precompR E'' : _) hf hg x₀ + have := hcg.convolutionExists_right (L.precompR E'' :) hf hg x₀ simp_rw [convolution_def, ContinuousLinearMap.integral_apply this] rfl @@ -1096,7 +1096,7 @@ theorem hasFDerivAt_convolution_right_with_param {g : P → G → E'} {s : Set P have I3 : AEStronglyMeasurable (fun a : G => (L (f a)).comp (g' (q₀.fst, q₀.snd - a))) μ := by have T : HasCompactSupport fun y => g' (q₀.1, y) := 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 + apply (HasCompactSupport.convolutionExists_right (L.precompR (P × G) :) T hf _ q₀.2).1 have : ContinuousOn g' (s ×ˢ univ) := hg.continuousOn_fderiv_of_isOpen (hs.prod isOpen_univ) le_rfl apply this.comp_continuous (continuous_const.prod_mk continuous_id') @@ -1206,7 +1206,7 @@ theorem contDiffOn_convolution_right_with_param_aux {G : Type uP} {E' : Type uP} filter_upwards [prod_mem_prod M1 M2] rintro ⟨p, y⟩ ⟨hp, hy⟩ exact hgs p y hp hy - apply ih (L.precompR (P × G) : _) B + apply ih (L.precompR (P × G) :) B convert hg.2.2 | htop ih => rw [contDiffOn_infty] at hg ⊢ diff --git a/Mathlib/Analysis/Distribution/SchwartzSpace.lean b/Mathlib/Analysis/Distribution/SchwartzSpace.lean index c5823ee573914..d259fdd46449b 100644 --- a/Mathlib/Analysis/Distribution/SchwartzSpace.lean +++ b/Mathlib/Analysis/Distribution/SchwartzSpace.lean @@ -185,7 +185,7 @@ variable [NormedField 𝕜] [NormedSpace 𝕜 F] [SMulCommClass ℝ 𝕜 F] theorem decay_smul_aux (k n : ℕ) (f : 𝓢(E, F)) (c : 𝕜) (x : E) : ‖x‖ ^ k * ‖iteratedFDeriv ℝ n (c • (f : E → F)) x‖ = ‖c‖ * ‖x‖ ^ k * ‖iteratedFDeriv ℝ n f x‖ := by - rw [mul_comm ‖c‖, mul_assoc, iteratedFDeriv_const_smul_apply (f.smooth _), + rw [mul_comm ‖c‖, mul_assoc, iteratedFDeriv_const_smul_apply (f.smooth _).contDiffAt, norm_smul c (iteratedFDeriv ℝ n (⇑f) x)] end Aux @@ -744,12 +744,15 @@ protected def evalCLM (m : E) : 𝓢(E, E →L[ℝ] F) →L[𝕜] 𝓢(E, F) := rintro ⟨k, n⟩ use {(k, n)}, ‖m‖, norm_nonneg _ intro f x - refine le_trans - (mul_le_mul_of_nonneg_left (norm_iteratedFDeriv_clm_apply_const f.2 (mod_cast le_top)) - (by positivity)) ?_ - move_mul [‖m‖] - gcongr ?_ * ‖m‖ - simp only [Finset.sup_singleton, schwartzSeminormFamily_apply, le_seminorm] + simp only [Finset.sup_singleton, schwartzSeminormFamily_apply] + calc + ‖x‖ ^ k * ‖iteratedFDeriv ℝ n (f · m) x‖ ≤ ‖x‖ ^ k * (‖m‖ * ‖iteratedFDeriv ℝ n f x‖) := by + gcongr + exact norm_iteratedFDeriv_clm_apply_const (f.smooth _).contDiffAt le_rfl + _ ≤ ‖m‖ * SchwartzMap.seminorm 𝕜 k n f := by + move_mul [‖m‖] + gcongr + apply le_seminorm end EvalCLM diff --git a/Mathlib/Analysis/Fourier/FourierTransform.lean b/Mathlib/Analysis/Fourier/FourierTransform.lean index a5d1ddb0067c9..c75680f9ddcd1 100644 --- a/Mathlib/Analysis/Fourier/FourierTransform.lean +++ b/Mathlib/Analysis/Fourier/FourierTransform.lean @@ -136,9 +136,6 @@ theorem fourierIntegral_convergent_iff (he : Continuous e) e.map_zero_eq_one, one_smul] at this -- the `(e _)` speeds up elaboration considerably exact this -@[deprecated (since := "2024-03-29")] -alias fourier_integral_convergent_iff := VectorFourier.fourierIntegral_convergent_iff - theorem fourierIntegral_add (he : Continuous e) (hL : Continuous fun p : V × W ↦ L p.1 p.2) {f g : V → E} (hf : Integrable f μ) (hg : Integrable g μ) : fourierIntegral e μ L (f + g) = fourierIntegral e μ L f + fourierIntegral e μ L g := by @@ -431,16 +428,11 @@ theorem fourierIntegral_real_eq (f : ℝ → E) (w : ℝ) : fourierIntegral f w = ∫ v : ℝ, 𝐞 (-(v * w)) • f v := rfl -@[deprecated (since := "2024-02-21")] alias fourierIntegral_def := fourierIntegral_real_eq - theorem fourierIntegral_real_eq_integral_exp_smul (f : ℝ → E) (w : ℝ) : 𝓕 f w = ∫ v : ℝ, Complex.exp (↑(-2 * π * v * w) * Complex.I) • f v := by simp_rw [fourierIntegral_real_eq, Circle.smul_def, Real.fourierChar_apply, mul_neg, neg_mul, mul_assoc] -@[deprecated (since := "2024-02-21")] -alias fourierIntegral_eq_integral_exp_smul := fourierIntegral_real_eq_integral_exp_smul - theorem fourierIntegral_continuousLinearMap_apply {F : Type*} [NormedAddCommGroup F] [NormedSpace ℝ F] {f : V → (F →L[ℝ] E)} {a : F} {v : V} (hf : Integrable f) : diff --git a/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean b/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean index 8543135aeea96..fe95597e697b1 100644 --- a/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean +++ b/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean @@ -200,7 +200,7 @@ lemma _root_.MeasureTheory.AEStronglyMeasurable.fourierSMulRight have aux1 : AEStronglyMeasurable (fun v ↦ (L v, f v)) μ := L.continuous.aestronglyMeasurable.prod_mk hf -- Elaboration without the expected type is faster here: - exact (aux0.comp_aestronglyMeasurable aux1 : _) + exact (aux0.comp_aestronglyMeasurable aux1 :) variable {f} @@ -386,8 +386,8 @@ lemma norm_iteratedFDeriv_fourierPowSMulRight -- second step: factor out the `(2 * π) ^ n` factor, and cancel it on both sides. have A : ContDiff ℝ K (fun y ↦ T (fun _ ↦ L y)) := (ContinuousMultilinearMap.contDiff _).comp (contDiff_pi.2 fun _ ↦ L.contDiff) - rw [iteratedFDeriv_const_smul_apply' (hf := (smulRightL ℝ (fun _ ↦ W) - E).isBoundedBilinearMap.contDiff.comp₂ (A.of_le hk) (hf.of_le hk)), + rw [iteratedFDeriv_const_smul_apply' (hf := ((smulRightL ℝ (fun _ ↦ W) + E).isBoundedBilinearMap.contDiff.comp₂ (A.of_le hk) (hf.of_le hk)).contDiffAt), norm_smul (β := V [×k]→L[ℝ] (W [×n]→L[ℝ] E))] simp only [norm_pow, norm_neg, norm_mul, RCLike.norm_ofNat, Complex.norm_eq_abs, abs_ofReal, _root_.abs_of_nonneg pi_nonneg, abs_I, mul_one, mul_assoc] diff --git a/Mathlib/Analysis/InnerProductSpace/Completion.lean b/Mathlib/Analysis/InnerProductSpace/Completion.lean index 330ff19ec4db2..3642a0acb90ca 100644 --- a/Mathlib/Analysis/InnerProductSpace/Completion.lean +++ b/Mathlib/Analysis/InnerProductSpace/Completion.lean @@ -86,7 +86,7 @@ protected theorem continuous_inner : protected theorem Continuous.inner {α : Type*} [TopologicalSpace α] {f g : α → Completion E} (hf : Continuous f) (hg : Continuous g) : Continuous (fun x : α => inner (f x) (g x) : α → 𝕜) := - UniformSpace.Completion.continuous_inner.comp (hf.prod_mk hg : _) + UniformSpace.Completion.continuous_inner.comp (hf.prod_mk hg :) instance innerProductSpace : InnerProductSpace 𝕜 (Completion E) where norm_sq_eq_inner x := diff --git a/Mathlib/Analysis/InnerProductSpace/JointEigenspace.lean b/Mathlib/Analysis/InnerProductSpace/JointEigenspace.lean index 9816edf6682a7..a340ee71187ce 100644 --- a/Mathlib/Analysis/InnerProductSpace/JointEigenspace.lean +++ b/Mathlib/Analysis/InnerProductSpace/JointEigenspace.lean @@ -111,6 +111,8 @@ theorem directSum_isInternal_of_commute (hA : A.IsSymmetric) (hB : B.IsSymmetric rw [Submodule.orthogonal_eq_bot_iff, iSup_prod, iSup_comm] exact iSup_iSup_eigenspace_inf_eigenspace_eq_top_of_commute hA hB hAB +open scoped Function -- required for scoped `on` notation + /-- A commuting family of symmetric linear maps on a finite dimensional inner product space is simultaneously diagonalizable. -/ theorem iSup_iInf_eq_top_of_commute {ι : Type*} {T : ι → E →ₗ[𝕜] E} diff --git a/Mathlib/Analysis/InnerProductSpace/LaxMilgram.lean b/Mathlib/Analysis/InnerProductSpace/LaxMilgram.lean index 9aa623e738fae..4e702fd38f691 100644 --- a/Mathlib/Analysis/InnerProductSpace/LaxMilgram.lean +++ b/Mathlib/Analysis/InnerProductSpace/LaxMilgram.lean @@ -76,7 +76,6 @@ theorem isClosed_range (coercive : IsCoercive B) : IsClosed (range B♯ : Set V) rcases coercive.antilipschitz with ⟨_, _, antilipschitz⟩ exact antilipschitz.isClosed_range B♯.uniformContinuous -@[deprecated (since := "2024-03-19")] alias closed_range := isClosed_range theorem range_eq_top (coercive : IsCoercive B) : range B♯ = ⊤ := by haveI := coercive.isClosed_range.completeSpace_coe diff --git a/Mathlib/Analysis/InnerProductSpace/Orthogonal.lean b/Mathlib/Analysis/InnerProductSpace/Orthogonal.lean index 70a253a0c3e16..876666070250b 100644 --- a/Mathlib/Analysis/InnerProductSpace/Orthogonal.lean +++ b/Mathlib/Analysis/InnerProductSpace/Orthogonal.lean @@ -350,6 +350,7 @@ theorem IsOrtho.comap_iff (f : E ≃ₗᵢ[𝕜] F) {U V : Submodule 𝕜 F} : U end Submodule +open scoped Function in -- required for scoped `on` notation theorem orthogonalFamily_iff_pairwise {ι} {V : ι → Submodule 𝕜 E} : (OrthogonalFamily 𝕜 (fun i => V i) fun i => (V i).subtypeₗᵢ) ↔ Pairwise ((· ⟂ ·) on V) := forall₃_congr fun _i _j _hij => diff --git a/Mathlib/Analysis/InnerProductSpace/Projection.lean b/Mathlib/Analysis/InnerProductSpace/Projection.lean index 902821ff1cdfd..2491cf9d734db 100644 --- a/Mathlib/Analysis/InnerProductSpace/Projection.lean +++ b/Mathlib/Analysis/InnerProductSpace/Projection.lean @@ -63,10 +63,10 @@ local notation "absR" => abs -- FIXME this monolithic proof causes a deterministic timeout with `-T50000` -- It should be broken in a sequence of more manageable pieces, -- perhaps with individual statements for the three steps below. -/-- Existence of minimizers +/-- **Existence of minimizers**, aka the **Hilbert projection theorem**. + Let `u` be a point in a real inner product space, and let `K` be a nonempty complete convex subset. -Then there exists a (unique) `v` in `K` that minimizes the distance `‖u - v‖` to `u`. - -/ +Then there exists a (unique) `v` in `K` that minimizes the distance `‖u - v‖` to `u`. -/ theorem exists_norm_eq_iInf_of_complete_convex {K : Set F} (ne : K.Nonempty) (h₁ : IsComplete K) (h₂ : Convex ℝ K) : ∀ u : F, ∃ v ∈ K, ‖u - v‖ = ⨅ w : K, ‖u - w‖ := fun u => by let δ := ⨅ w : K, ‖u - w‖ @@ -1288,7 +1288,7 @@ abbrev OrthogonalFamily.decomposition [DecidableEq ι] [Fintype ι] {V : ι → -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 erw [DFinsupp.sumAddHom_apply]; rw [DFinsupp.sum_eq_sum_fintype] · simp_rw [Equiv.apply_symm_apply, AddSubmonoidClass.coe_subtype] - exact hV.sum_projection_of_mem_iSup _ ((h.ge : _) Submodule.mem_top) + exact hV.sum_projection_of_mem_iSup _ ((h.ge :) Submodule.mem_top) · intro i exact map_zero _ right_inv x := by diff --git a/Mathlib/Analysis/InnerProductSpace/TwoDim.lean b/Mathlib/Analysis/InnerProductSpace/TwoDim.lean index a9afb01f2a21b..ce34bcdfe2375 100644 --- a/Mathlib/Analysis/InnerProductSpace/TwoDim.lean +++ b/Mathlib/Analysis/InnerProductSpace/TwoDim.lean @@ -79,10 +79,6 @@ lemma FiniteDimensional.of_fact_finrank_eq_two {K V : Type*} [DivisionRing K] attribute [local instance] FiniteDimensional.of_fact_finrank_eq_two -@[deprecated (since := "2024-02-02")] -alias FiniteDimensional.finiteDimensional_of_fact_finrank_eq_two := - FiniteDimensional.of_fact_finrank_eq_two - variable {E : Type*} [NormedAddCommGroup E] [InnerProductSpace ℝ E] [Fact (finrank ℝ E = 2)] (o : Orientation ℝ E (Fin 2)) diff --git a/Mathlib/Analysis/LocallyConvex/Basic.lean b/Mathlib/Analysis/LocallyConvex/Basic.lean index ffe0e4d997485..f6d8feaab62d1 100644 --- a/Mathlib/Analysis/LocallyConvex/Basic.lean +++ b/Mathlib/Analysis/LocallyConvex/Basic.lean @@ -215,8 +215,6 @@ theorem Balanced.absorbs_self (hA : Balanced 𝕜 A) : Absorbs 𝕜 A A := theorem Balanced.smul_mem_iff (hs : Balanced 𝕜 s) (h : ‖a‖ = ‖b‖) : a • x ∈ s ↔ b • x ∈ s := ⟨(hs.smul_mem_mono · h.ge), (hs.smul_mem_mono · h.le)⟩ -@[deprecated (since := "2024-02-02")] alias Balanced.mem_smul_iff := Balanced.smul_mem_iff - variable [TopologicalSpace E] [ContinuousSMul 𝕜 E] /-- Every neighbourhood of the origin is absorbent. -/ @@ -235,10 +233,6 @@ theorem Balanced.zero_insert_interior (hA : Balanced 𝕜 A) : apply insert_subset_insert exact ((isOpenMap_smul₀ h).mapsTo_interior <| hA.smul_mem ha).image_subset -@[deprecated Balanced.zero_insert_interior (since := "2024-02-03")] -theorem balanced_zero_union_interior (hA : Balanced 𝕜 A) : Balanced 𝕜 ((0 : Set E) ∪ interior A) := - hA.zero_insert_interior - /-- The interior of a balanced set is balanced if it contains the origin. -/ protected theorem Balanced.interior (hA : Balanced 𝕜 A) (h : (0 : E) ∈ interior A) : Balanced 𝕜 (interior A) := by @@ -255,9 +249,6 @@ section NontriviallyNormedField variable [NontriviallyNormedField 𝕜] [AddCommGroup E] [Module 𝕜 E] {s : Set E} -@[deprecated Absorbent.zero_mem (since := "2024-02-02")] -theorem Absorbent.zero_mem' (hs : Absorbent 𝕜 s) : (0 : E) ∈ s := hs.zero_mem - variable [Module ℝ E] [SMulCommClass ℝ 𝕜 E] protected theorem Balanced.convexHull (hs : Balanced 𝕜 s) : Balanced 𝕜 (convexHull ℝ s) := by @@ -269,8 +260,6 @@ protected theorem Balanced.convexHull (hs : Balanced 𝕜 s) : Balanced 𝕜 (co simp only [smul_add, ← smul_comm] exact convex_convexHull ℝ s (hx a ha) (hy a ha) hu hv huv -@[deprecated (since := "2024-02-02")] alias balanced_convexHull_of_balanced := Balanced.convexHull - end NontriviallyNormedField section Real diff --git a/Mathlib/Analysis/Matrix.lean b/Mathlib/Analysis/Matrix.lean index 7a1f22fb8be30..0101a1bd7d633 100644 --- a/Mathlib/Analysis/Matrix.lean +++ b/Mathlib/Analysis/Matrix.lean @@ -108,7 +108,7 @@ theorem nnnorm_map_eq (A : Matrix m n α) (f : α → β) (hf : ∀ a, ‖f a‖ @[simp] theorem norm_map_eq (A : Matrix m n α) (f : α → β) (hf : ∀ a, ‖f a‖ = ‖a‖) : ‖A.map f‖ = ‖A‖ := - (congr_arg ((↑) : ℝ≥0 → ℝ) <| nnnorm_map_eq A f fun a => Subtype.ext <| hf a : _) + (congr_arg ((↑) : ℝ≥0 → ℝ) <| nnnorm_map_eq A f fun a => Subtype.ext <| hf a :) @[simp] theorem nnnorm_transpose (A : Matrix m n α) : ‖Aᵀ‖₊ = ‖A‖₊ := @@ -248,39 +248,27 @@ theorem linfty_opNorm_def (A : Matrix m n α) : change ‖fun i => (WithLp.equiv 1 _).symm (A i)‖ = _ simp [Pi.norm_def, PiLp.nnnorm_eq_of_L1] -@[deprecated (since := "2024-02-02")] alias linfty_op_norm_def := linfty_opNorm_def - theorem linfty_opNNNorm_def (A : Matrix m n α) : ‖A‖₊ = (Finset.univ : Finset m).sup fun i : m => ∑ j : n, ‖A i j‖₊ := Subtype.ext <| linfty_opNorm_def A -@[deprecated (since := "2024-02-02")] alias linfty_op_nnnorm_def := linfty_opNNNorm_def - @[simp] theorem linfty_opNNNorm_col (v : m → α) : ‖col ι v‖₊ = ‖v‖₊ := by rw [linfty_opNNNorm_def, Pi.nnnorm_def] simp -@[deprecated (since := "2024-02-02")] alias linfty_op_nnnorm_col := linfty_opNNNorm_col - @[simp] theorem linfty_opNorm_col (v : m → α) : ‖col ι v‖ = ‖v‖ := congr_arg ((↑) : ℝ≥0 → ℝ) <| linfty_opNNNorm_col v -@[deprecated (since := "2024-02-02")] alias linfty_op_norm_col := linfty_opNorm_col - @[simp] theorem linfty_opNNNorm_row (v : n → α) : ‖row ι v‖₊ = ∑ i, ‖v i‖₊ := by simp [linfty_opNNNorm_def] -@[deprecated (since := "2024-02-02")] alias linfty_op_nnnorm_row := linfty_opNNNorm_row - @[simp] theorem linfty_opNorm_row (v : n → α) : ‖row ι v‖ = ∑ i, ‖v i‖ := (congr_arg ((↑) : ℝ≥0 → ℝ) <| linfty_opNNNorm_row v).trans <| by simp [NNReal.coe_sum] -@[deprecated (since := "2024-02-02")] alias linfty_op_norm_row := linfty_opNorm_row - @[simp] theorem linfty_opNNNorm_diagonal [DecidableEq m] (v : m → α) : ‖diagonal v‖₊ = ‖v‖₊ := by rw [linfty_opNNNorm_def, Pi.nnnorm_def] @@ -289,14 +277,10 @@ theorem linfty_opNNNorm_diagonal [DecidableEq m] (v : m → α) : ‖diagonal v · rw [diagonal_apply_ne' _ hij, nnnorm_zero] · rw [diagonal_apply_eq] -@[deprecated (since := "2024-02-02")] alias linfty_op_nnnorm_diagonal := linfty_opNNNorm_diagonal - @[simp] theorem linfty_opNorm_diagonal [DecidableEq m] (v : m → α) : ‖diagonal v‖ = ‖v‖ := congr_arg ((↑) : ℝ≥0 → ℝ) <| linfty_opNNNorm_diagonal v -@[deprecated (since := "2024-02-02")] alias linfty_op_norm_diagonal := linfty_opNorm_diagonal - end SeminormedAddCommGroup section NonUnitalSeminormedRing @@ -320,24 +304,16 @@ theorem linfty_opNNNorm_mul (A : Matrix l m α) (B : Matrix m n α) : ‖A * B simp_rw [← Finset.sum_mul, ← NNReal.finset_sup_mul] rfl -@[deprecated (since := "2024-02-02")] alias linfty_op_nnnorm_mul := linfty_opNNNorm_mul - theorem linfty_opNorm_mul (A : Matrix l m α) (B : Matrix m n α) : ‖A * B‖ ≤ ‖A‖ * ‖B‖ := linfty_opNNNorm_mul _ _ -@[deprecated (since := "2024-02-02")] alias linfty_op_norm_mul := linfty_opNorm_mul - theorem linfty_opNNNorm_mulVec (A : Matrix l m α) (v : m → α) : ‖A *ᵥ v‖₊ ≤ ‖A‖₊ * ‖v‖₊ := by rw [← linfty_opNNNorm_col (ι := Fin 1) (A *ᵥ v), ← linfty_opNNNorm_col v (ι := Fin 1)] exact linfty_opNNNorm_mul A (col (Fin 1) v) -@[deprecated (since := "2024-02-02")] alias linfty_op_nnnorm_mulVec := linfty_opNNNorm_mulVec - theorem linfty_opNorm_mulVec (A : Matrix l m α) (v : m → α) : ‖A *ᵥ v‖ ≤ ‖A‖ * ‖v‖ := linfty_opNNNorm_mulVec _ _ -@[deprecated (since := "2024-02-02")] alias linfty_op_norm_mulVec := linfty_opNorm_mulVec - end NonUnitalSeminormedRing /-- Seminormed non-unital ring instance (using sup norm of L1 norm) for matrices over a semi normed @@ -435,15 +411,10 @@ lemma linfty_opNNNorm_eq_opNNNorm (A : Matrix m n α) : nnnorm_one, mul_one] at hN exact hN -@[deprecated (since := "2024-02-02")] -alias linfty_op_nnnorm_eq_op_nnnorm := linfty_opNNNorm_eq_opNNNorm - lemma linfty_opNorm_eq_opNorm (A : Matrix m n α) : ‖A‖ = ‖ContinuousLinearMap.mk (Matrix.mulVecLin A)‖ := congr_arg NNReal.toReal (linfty_opNNNorm_eq_opNNNorm A) -@[deprecated (since := "2024-02-02")] alias linfty_op_norm_eq_op_norm := linfty_opNorm_eq_opNorm - variable [DecidableEq n] @[simp] lemma linfty_opNNNorm_toMatrix (f : (n → α) →L[α] (m → α)) : @@ -451,14 +422,10 @@ variable [DecidableEq n] rw [linfty_opNNNorm_eq_opNNNorm] simp only [← toLin'_apply', toLin'_toMatrix'] -@[deprecated (since := "2024-02-02")] alias linfty_op_nnnorm_toMatrix := linfty_opNNNorm_toMatrix - @[simp] lemma linfty_opNorm_toMatrix (f : (n → α) →L[α] (m → α)) : ‖LinearMap.toMatrix' (↑f : (n → α) →ₗ[α] (m → α))‖ = ‖f‖ := congr_arg NNReal.toReal (linfty_opNNNorm_toMatrix f) -@[deprecated (since := "2024-02-02")] alias linfty_op_norm_toMatrix := linfty_opNorm_toMatrix - end end LinftyOp @@ -526,7 +493,7 @@ theorem frobenius_nnnorm_map_eq (A : Matrix m n α) (f : α → β) (hf : ∀ a, @[simp] theorem frobenius_norm_map_eq (A : Matrix m n α) (f : α → β) (hf : ∀ a, ‖f a‖ = ‖a‖) : ‖A.map f‖ = ‖A‖ := - (congr_arg ((↑) : ℝ≥0 → ℝ) <| frobenius_nnnorm_map_eq A f fun a => Subtype.ext <| hf a : _) + (congr_arg ((↑) : ℝ≥0 → ℝ) <| frobenius_nnnorm_map_eq A f fun a => Subtype.ext <| hf a :) @[simp] theorem frobenius_nnnorm_transpose (A : Matrix m n α) : ‖Aᵀ‖₊ = ‖A‖₊ := by @@ -586,7 +553,7 @@ theorem frobenius_nnnorm_diagonal [DecidableEq n] (v : n → α) : @[simp] theorem frobenius_norm_diagonal [DecidableEq n] (v : n → α) : ‖diagonal v‖ = ‖(WithLp.equiv 2 _).symm v‖ := - (congr_arg ((↑) : ℝ≥0 → ℝ) <| frobenius_nnnorm_diagonal v : _).trans rfl + (congr_arg ((↑) : ℝ≥0 → ℝ) <| frobenius_nnnorm_diagonal v :).trans rfl end SeminormedAddCommGroup diff --git a/Mathlib/Analysis/Normed/Algebra/Basic.lean b/Mathlib/Analysis/Normed/Algebra/Basic.lean index d44b75699bb05..6da32173a5821 100644 --- a/Mathlib/Analysis/Normed/Algebra/Basic.lean +++ b/Mathlib/Analysis/Normed/Algebra/Basic.lean @@ -44,7 +44,7 @@ instance [ProperSpace 𝕜] : CompactSpace (characterSpace 𝕜 A) := by have h : characterSpace 𝕜 A ⊆ toNormedDual ⁻¹' Metric.closedBall 0 ‖(1 : A)‖ := by intro φ hφ rw [Set.mem_preimage, mem_closedBall_zero_iff] - exact (norm_le_norm_one ⟨φ, ⟨hφ.1, hφ.2⟩⟩ : _) + exact (norm_le_norm_one ⟨φ, ⟨hφ.1, hφ.2⟩⟩ :) exact (isCompact_closedBall 𝕜 0 _).of_isClosed_subset CharacterSpace.isClosed h end CharacterSpace diff --git a/Mathlib/Analysis/Normed/Algebra/Exponential.lean b/Mathlib/Analysis/Normed/Algebra/Exponential.lean index 2c3861688440d..9166797fa69d8 100644 --- a/Mathlib/Analysis/Normed/Algebra/Exponential.lean +++ b/Mathlib/Analysis/Normed/Algebra/Exponential.lean @@ -474,6 +474,7 @@ theorem exp_mem_unitary_of_mem_skewAdjoint [StarRing 𝔸] [ContinuousStar 𝔸] end +open scoped Function in -- required for scoped `on` notation /-- In a Banach-algebra `𝔸` over `𝕂 = ℝ` or `𝕂 = ℂ`, if a family of elements `f i` mutually commute then `NormedSpace.exp 𝕂 (∑ i, f i) = ∏ i, NormedSpace.exp 𝕂 (f i)`. -/ theorem exp_sum_of_commute {ι} (s : Finset ι) (f : ι → 𝔸) diff --git a/Mathlib/Analysis/Normed/Algebra/MatrixExponential.lean b/Mathlib/Analysis/Normed/Algebra/MatrixExponential.lean index fe33c719a4832..3bbd2707d64b9 100644 --- a/Mathlib/Analysis/Normed/Algebra/MatrixExponential.lean +++ b/Mathlib/Analysis/Normed/Algebra/MatrixExponential.lean @@ -124,6 +124,7 @@ nonrec theorem exp_add_of_commute (A B : Matrix m m 𝔸) (h : Commute A B) : letI : NormedAlgebra 𝕂 (Matrix m m 𝔸) := Matrix.linftyOpNormedAlgebra exact exp_add_of_commute h +open scoped Function in -- required for scoped `on` notation nonrec theorem exp_sum_of_commute {ι} (s : Finset ι) (f : ι → Matrix m m 𝔸) (h : (s : Set ι).Pairwise (Commute on f)) : exp 𝕂 (∑ i ∈ s, f i) = diff --git a/Mathlib/Analysis/Normed/Field/Basic.lean b/Mathlib/Analysis/Normed/Field/Basic.lean index 723a44e211ecc..9d8514ced1889 100644 --- a/Mathlib/Analysis/Normed/Field/Basic.lean +++ b/Mathlib/Analysis/Normed/Field/Basic.lean @@ -292,7 +292,7 @@ instance (priority := 75) NonUnitalSubalgebraClass.nonUnitalNormedRing {S 𝕜 E instance ULift.nonUnitalSeminormedRing : NonUnitalSeminormedRing (ULift α) := { ULift.seminormedAddCommGroup, ULift.nonUnitalRing with - norm_mul := fun x y => (norm_mul_le x.down y.down : _) } + norm_mul := fun x y => (norm_mul_le x.down y.down :) } /-- Non-unital seminormed ring structure on the product of two non-unital seminormed rings, using the sup norm. -/ diff --git a/Mathlib/Analysis/Normed/Field/InfiniteSum.lean b/Mathlib/Analysis/Normed/Field/InfiniteSum.lean index 195aa8c9adc27..faf7d5a691a90 100644 --- a/Mathlib/Analysis/Normed/Field/InfiniteSum.lean +++ b/Mathlib/Analysis/Normed/Field/InfiniteSum.lean @@ -36,7 +36,7 @@ theorem Summable.mul_norm {f : ι → R} {g : ι' → R} (hf : Summable fun x => (hg : Summable fun x => ‖g x‖) : Summable fun x : ι × ι' => ‖f x.1 * g x.2‖ := .of_nonneg_of_le (fun _ ↦ norm_nonneg _) (fun x => norm_mul_le (f x.1) (g x.2)) - (hf.mul_of_nonneg hg (fun x => norm_nonneg <| f x) fun x => norm_nonneg <| g x : _) + (hf.mul_of_nonneg hg (fun x => norm_nonneg <| f x) fun x => norm_nonneg <| g x :) theorem summable_mul_of_summable_norm [CompleteSpace R] {f : ι → R} {g : ι' → R} (hf : Summable fun x => ‖f x‖) (hg : Summable fun x => ‖g x‖) : diff --git a/Mathlib/Analysis/Normed/Group/AddTorsor.lean b/Mathlib/Analysis/Normed/Group/AddTorsor.lean index e3a3f8109709e..d9089e512a53f 100644 --- a/Mathlib/Analysis/Normed/Group/AddTorsor.lean +++ b/Mathlib/Analysis/Normed/Group/AddTorsor.lean @@ -235,7 +235,7 @@ variable [TopologicalSpace α] theorem Continuous.vsub {f g : α → P} (hf : Continuous f) (hg : Continuous g) : Continuous (f -ᵥ g) := - continuous_vsub.comp (hf.prod_mk hg : _) + continuous_vsub.comp (hf.prod_mk hg :) nonrec theorem ContinuousAt.vsub {f g : α → P} {x : α} (hf : ContinuousAt f x) (hg : ContinuousAt g x) : diff --git a/Mathlib/Analysis/Normed/Group/Basic.lean b/Mathlib/Analysis/Normed/Group/Basic.lean index fc2858dbb2006..493bd417df034 100644 --- a/Mathlib/Analysis/Normed/Group/Basic.lean +++ b/Mathlib/Analysis/Normed/Group/Basic.lean @@ -1240,10 +1240,10 @@ theorem le_norm_self (r : ℝ) : r ≤ ‖r‖ := @[deprecated (since := "2024-04-05")] alias nnnorm_coe_nat := nnnorm_natCast @[simp 1100] lemma norm_ofNat (n : ℕ) [n.AtLeastTwo] : - ‖(no_index (OfNat.ofNat n) : ℝ)‖ = OfNat.ofNat n := norm_natCast n + ‖(ofNat(n) : ℝ)‖ = OfNat.ofNat n := norm_natCast n @[simp 1100] lemma nnnorm_ofNat (n : ℕ) [n.AtLeastTwo] : - ‖(no_index (OfNat.ofNat n) : ℝ)‖₊ = OfNat.ofNat n := nnnorm_natCast n + ‖(ofNat(n) : ℝ)‖₊ = OfNat.ofNat n := nnnorm_natCast n lemma norm_two : ‖(2 : ℝ)‖ = 2 := abs_of_pos zero_lt_two lemma nnnorm_two : ‖(2 : ℝ)‖₊ = 2 := NNReal.eq <| by simp diff --git a/Mathlib/Analysis/Normed/Group/Constructions.lean b/Mathlib/Analysis/Normed/Group/Constructions.lean index 95eb19382e965..5a2a7e261e62b 100644 --- a/Mathlib/Analysis/Normed/Group/Constructions.lean +++ b/Mathlib/Analysis/Normed/Group/Constructions.lean @@ -229,6 +229,8 @@ instance Prod.toNorm : Norm (E × F) where norm x := ‖x.1‖ ⊔ ‖x.2‖ lemma Prod.norm_def (x : E × F) : ‖x‖ = max ‖x.1‖ ‖x.2‖ := rfl +@[simp] lemma Prod.norm_mk (x : E) (y : F) : ‖(x, y)‖ = max ‖x‖ ‖y‖ := rfl + lemma norm_fst_le (x : E × F) : ‖x.1‖ ≤ ‖x‖ := le_max_left _ _ lemma norm_snd_le (x : E × F) : ‖x.2‖ ≤ ‖x‖ := le_max_right _ _ @@ -246,8 +248,16 @@ instance Prod.seminormedGroup : SeminormedGroup (E × F) where dist_eq x y := by simp only [Prod.norm_def, Prod.dist_eq, dist_eq_norm_div, Prod.fst_div, Prod.snd_div] -@[to_additive Prod.nnnorm_def'] -lemma Prod.nnorm_def (x : E × F) : ‖x‖₊ = max ‖x.1‖₊ ‖x.2‖₊ := rfl +/-- Multiplicative version of `Prod.nnnorm_def`. +Earlier, this names was used for the additive version. -/ +@[to_additive Prod.nnnorm_def] +lemma Prod.nnnorm_def' (x : E × F) : ‖x‖₊ = max ‖x.1‖₊ ‖x.2‖₊ := rfl + +@[deprecated (since := "2025-01-02")] alias Prod.nnorm_def := Prod.nnnorm_def' + +/-- Multiplicative version of `Prod.nnnorm_mk`. -/ +@[to_additive (attr := simp) Prod.nnnorm_mk] +lemma Prod.nnnorm_mk' (x : E) (y : F) : ‖(x, y)‖₊ = max ‖x‖₊ ‖y‖₊ := rfl end SeminormedGroup diff --git a/Mathlib/Analysis/Normed/Group/SemiNormedGrp/Kernels.lean b/Mathlib/Analysis/Normed/Group/SemiNormedGrp/Kernels.lean index 302ed81b07f5c..ba8df90a57055 100644 --- a/Mathlib/Analysis/Normed/Group/SemiNormedGrp/Kernels.lean +++ b/Mathlib/Analysis/Normed/Group/SemiNormedGrp/Kernels.lean @@ -80,7 +80,7 @@ instance : HasCokernels SemiNormedGrp₁.{u} where erw [zero_apply]) fun _ _ w => Subtype.eq - (NormedAddGroupHom.lift_unique f.1.range _ _ _ (congr_arg Subtype.val w : _)) } + (NormedAddGroupHom.lift_unique f.1.range _ _ _ (congr_arg Subtype.val w :)) } -- Sanity check example : HasCokernels SemiNormedGrp₁ := by infer_instance diff --git a/Mathlib/Analysis/Normed/Lp/LpEquiv.lean b/Mathlib/Analysis/Normed/Lp/LpEquiv.lean index d77bc7fea6e3e..bd8818ac7ee46 100644 --- a/Mathlib/Analysis/Normed/Lp/LpEquiv.lean +++ b/Mathlib/Analysis/Normed/Lp/LpEquiv.lean @@ -133,7 +133,6 @@ noncomputable def AddEquiv.lpBCF : lp (fun _ : α ↦ E) ∞ ≃+ (α →ᵇ E) right_inv _f := rfl map_add' _f _g := rfl -@[deprecated (since := "2024-03-16")] alias AddEquiv.lpBcf := AddEquiv.lpBCF theorem coe_addEquiv_lpBCF (f : lp (fun _ : α ↦ E) ∞) : (AddEquiv.lpBCF f : α → E) = f := rfl @@ -149,7 +148,6 @@ noncomputable def lpBCFₗᵢ : lp (fun _ : α ↦ E) ∞ ≃ₗᵢ[𝕜] α → map_smul' := fun _ _ ↦ rfl norm_map' := fun f ↦ by simp only [norm_eq_iSup_norm, lp.norm_eq_ciSup]; rfl } -@[deprecated (since := "2024-03-16")] alias lpBcfₗᵢ := lpBCFₗᵢ variable {𝕜 E} @@ -168,7 +166,6 @@ noncomputable def RingEquiv.lpBCF : lp (fun _ : α ↦ R) ∞ ≃+* (α →ᵇ R { @AddEquiv.lpBCF _ R _ _ _ with map_mul' := fun _f _g => rfl } -@[deprecated (since := "2024-03-16")] alias RingEquiv.lpBcf := RingEquiv.lpBCF variable {R} @@ -187,7 +184,6 @@ variable (α) noncomputable def AlgEquiv.lpBCF : lp (fun _ : α ↦ A) ∞ ≃ₐ[𝕜] α →ᵇ A := { RingEquiv.lpBCF A with commutes' := fun _k ↦ rfl } -@[deprecated (since := "2024-03-16")] alias AlgEquiv.lpBcf := AlgEquiv.lpBCF variable {α A 𝕜} diff --git a/Mathlib/Analysis/Normed/Lp/ProdLp.lean b/Mathlib/Analysis/Normed/Lp/ProdLp.lean index d94faf96a1d7e..be6e704e2ad40 100644 --- a/Mathlib/Analysis/Normed/Lp/ProdLp.lean +++ b/Mathlib/Analysis/Normed/Lp/ProdLp.lean @@ -616,7 +616,7 @@ theorem prod_nnnorm_eq_sup (f : WithLp ∞ (α × β)) : ‖f‖₊ = ‖f.fst norm_cast @[simp] theorem prod_nnnorm_equiv (f : WithLp ∞ (α × β)) : ‖WithLp.equiv ⊤ _ f‖₊ = ‖f‖₊ := by - rw [prod_nnnorm_eq_sup, Prod.nnnorm_def', equiv_fst, equiv_snd] + rw [prod_nnnorm_eq_sup, Prod.nnnorm_def, equiv_fst, equiv_snd] @[simp] theorem prod_nnnorm_equiv_symm (f : α × β) : ‖(WithLp.equiv ⊤ _).symm f‖₊ = ‖f‖₊ := (prod_nnnorm_equiv _).symm diff --git a/Mathlib/Analysis/Normed/Module/Basic.lean b/Mathlib/Analysis/Normed/Module/Basic.lean index 615129ab6a77a..3ca89f37a1d43 100644 --- a/Mathlib/Analysis/Normed/Module/Basic.lean +++ b/Mathlib/Analysis/Normed/Module/Basic.lean @@ -102,7 +102,7 @@ open NormedField instance ULift.normedSpace : NormedSpace 𝕜 (ULift E) := { __ := ULift.seminormedAddCommGroup (E := E), __ := ULift.module' - norm_smul_le := fun s x => (norm_smul_le s x.down : _) } + norm_smul_le := fun s x => (norm_smul_le s x.down :) } /-- The product of two normed spaces is a normed space, with the sup norm. -/ instance Prod.normedSpace : NormedSpace 𝕜 (E × F) := diff --git a/Mathlib/Analysis/Normed/Module/FiniteDimension.lean b/Mathlib/Analysis/Normed/Module/FiniteDimension.lean index 789c690f2042f..90fc73c10bcf6 100644 --- a/Mathlib/Analysis/Normed/Module/FiniteDimension.lean +++ b/Mathlib/Analysis/Normed/Module/FiniteDimension.lean @@ -287,15 +287,11 @@ theorem Basis.opNNNorm_le {ι : Type*} [Fintype ι] (v : Basis ι 𝕜 E) {u : E _ ≤ Fintype.card ι • (‖φ‖₊ * ‖e‖₊) := nsmul_le_nsmul_right (φ.le_opNNNorm e) _ _ = Fintype.card ι • ‖φ‖₊ * M * ‖e‖₊ := by simp only [smul_mul_assoc, mul_right_comm] -@[deprecated (since := "2024-02-02")] alias Basis.op_nnnorm_le := Basis.opNNNorm_le - theorem Basis.opNorm_le {ι : Type*} [Fintype ι] (v : Basis ι 𝕜 E) {u : E →L[𝕜] F} {M : ℝ} (hM : 0 ≤ M) (hu : ∀ i, ‖u (v i)‖ ≤ M) : ‖u‖ ≤ Fintype.card ι • ‖v.equivFunL.toContinuousLinearMap‖ * M := by simpa using NNReal.coe_le_coe.mpr (v.opNNNorm_le ⟨M, hM⟩ hu) -@[deprecated (since := "2024-02-02")] alias Basis.op_norm_le := Basis.opNorm_le - /-- A weaker version of `Basis.opNNNorm_le` that abstracts away the value of `C`. -/ theorem Basis.exists_opNNNorm_le {ι : Type*} [Finite ι] (v : Basis ι 𝕜 E) : ∃ C > (0 : ℝ≥0), ∀ {u : E →L[𝕜] F} (M : ℝ≥0), (∀ i, ‖u (v i)‖₊ ≤ M) → ‖u‖₊ ≤ C * M := by @@ -305,8 +301,6 @@ theorem Basis.exists_opNNNorm_le {ι : Type*} [Finite ι] (v : Basis ι 𝕜 E) zero_lt_one.trans_le (le_max_right _ _), fun {u} M hu => (v.opNNNorm_le M hu).trans <| mul_le_mul_of_nonneg_right (le_max_left _ _) (zero_le M)⟩ -@[deprecated (since := "2024-02-02")] alias Basis.exists_op_nnnorm_le := Basis.exists_opNNNorm_le - /-- A weaker version of `Basis.opNorm_le` that abstracts away the value of `C`. -/ theorem Basis.exists_opNorm_le {ι : Type*} [Finite ι] (v : Basis ι 𝕜 E) : ∃ C > (0 : ℝ), ∀ {u : E →L[𝕜] F} {M : ℝ}, 0 ≤ M → (∀ i, ‖u (v i)‖ ≤ M) → ‖u‖ ≤ C * M := by @@ -316,8 +310,6 @@ theorem Basis.exists_opNorm_le {ι : Type*} [Finite ι] (v : Basis ι 𝕜 E) : intro u M hM H simpa using h ⟨M, hM⟩ H -@[deprecated (since := "2024-02-02")] alias Basis.exists_op_norm_le := Basis.exists_opNorm_le - instance [FiniteDimensional 𝕜 E] [SecondCountableTopology F] : SecondCountableTopology (E →L[𝕜] F) := by set d := Module.finrank 𝕜 E @@ -455,27 +447,18 @@ theorem FiniteDimensional.of_isCompact_closedBall₀ {r : ℝ} (rpos : 0 < r) exact φmono (Nat.lt_succ_self N) _ < ‖c‖ := hN (N + 1) (Nat.le_succ N) -@[deprecated (since := "2024-02-02")] -alias finiteDimensional_of_isCompact_closedBall₀ := FiniteDimensional.of_isCompact_closedBall₀ - /-- **Riesz's theorem**: if a closed ball of positive radius is compact in a vector space, then the space is finite-dimensional. -/ theorem FiniteDimensional.of_isCompact_closedBall {r : ℝ} (rpos : 0 < r) {c : E} (h : IsCompact (Metric.closedBall c r)) : FiniteDimensional 𝕜 E := .of_isCompact_closedBall₀ 𝕜 rpos <| by simpa using h.vadd (-c) -@[deprecated (since := "2024-02-02")] -alias finiteDimensional_of_isCompact_closedBall := FiniteDimensional.of_isCompact_closedBall - /-- **Riesz's theorem**: a locally compact normed vector space is finite-dimensional. -/ theorem FiniteDimensional.of_locallyCompactSpace [LocallyCompactSpace E] : FiniteDimensional 𝕜 E := let ⟨_r, rpos, hr⟩ := exists_isCompact_closedBall (0 : E) .of_isCompact_closedBall₀ 𝕜 rpos hr -@[deprecated (since := "2024-02-02")] -alias finiteDimensional_of_locallyCompactSpace := FiniteDimensional.of_locallyCompactSpace - /-- If a function has compact support, then either the function is trivial or the space is finite-dimensional. -/ theorem HasCompactSupport.eq_zero_or_finiteDimensional {X : Type*} [TopologicalSpace X] [Zero X] @@ -655,6 +638,11 @@ theorem summable_norm_iff {α E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ alias ⟨_, Summable.norm⟩ := summable_norm_iff +theorem summable_of_sum_range_norm_le {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] + [FiniteDimensional ℝ E] {c : ℝ} {f : ℕ → E} (h : ∀ n, ∑ i ∈ Finset.range n, ‖f i‖ ≤ c) : + Summable f := + summable_norm_iff.mp <| summable_of_sum_range_le (fun _ ↦ norm_nonneg _) h + theorem summable_of_isBigO' {ι E F : Type*} [NormedAddCommGroup E] [CompleteSpace E] [NormedAddCommGroup F] [NormedSpace ℝ F] [FiniteDimensional ℝ F] {f : ι → E} {g : ι → F} (hg : Summable g) (h : f =O[cofinite] g) : Summable f := diff --git a/Mathlib/Analysis/Normed/Ring/Seminorm.lean b/Mathlib/Analysis/Normed/Ring/Seminorm.lean index 46f441641f167..7856ea39c83fa 100644 --- a/Mathlib/Analysis/Normed/Ring/Seminorm.lean +++ b/Mathlib/Analysis/Normed/Ring/Seminorm.lean @@ -23,6 +23,8 @@ For a ring `R`: * `MulRingSeminorm`: A multiplicative seminorm on a ring `R` is a ring seminorm that preserves multiplication. * `MulRingNorm`: A multiplicative norm on a ring `R` is a ring norm that preserves multiplication. + `MulRingNorm R` is essentially the same as `AbsoluteValue R ℝ`, and it is recommended to + use the latter instead to avoid duplicating results. ## Notes @@ -60,7 +62,10 @@ structure MulRingSeminorm (R : Type*) [NonAssocRing R] extends AddGroupSeminorm MonoidWithZeroHom R ℝ /-- A multiplicative norm on a ring `R` is a multiplicative ring seminorm such that `f x = 0` -implies `x = 0`. -/ +implies `x = 0`. + +It is recommended to use `AbsoluteValue R ℝ` instead (which works for `Semiring R` +and is equivalent to `MulRingNorm R` for a nontrivial `Ring R`). -/ structure MulRingNorm (R : Type*) [NonAssocRing R] extends MulRingSeminorm R, AddGroupNorm R attribute [nolint docBlame] @@ -351,20 +356,56 @@ theorem apply_one (x : R) : (1 : MulRingNorm R) x = if x = 0 then 0 else 1 := instance : Inhabited (MulRingNorm R) := ⟨1⟩ +section MulRingNorm_equiv_AbsoluteValue + +variable {R : Type*} [Ring R] [Nontrivial R] + +/-- The equivalence of `MulRingNorm R` and `AbsoluteValue R ℝ` when `R` is a nontrivial ring. -/ +def mulRingNormEquivAbsoluteValue : MulRingNorm R ≃ AbsoluteValue R ℝ where + toFun N := { + toFun := N.toFun + map_mul' := N.map_mul' + nonneg' := apply_nonneg N + eq_zero' x := ⟨N.eq_zero_of_map_eq_zero' x, fun h ↦ h ▸ N.map_zero'⟩ + add_le' := N.add_le' + } + invFun v := { + toFun := v.toFun + map_zero' := (v.eq_zero' 0).mpr rfl + add_le' := v.add_le' + neg' := v.map_neg + map_one' := v.map_one + map_mul' := v.map_mul' + eq_zero_of_map_eq_zero' x := (v.eq_zero' x).mp + } + left_inv N := by ext1 x; simp only [MulRingSeminorm.toFun_eq_coe] -- `simp` does not work + right_inv v := by ext1 x; simp + +lemma mulRingNormEquivAbsoluteValue_apply (N : MulRingNorm R) (x : R) : + mulRingNormEquivAbsoluteValue N x = N x := rfl + +lemma mulRingNormEquivAbsoluteValue_symm_apply (v : AbsoluteValue R ℝ) (x : R) : + mulRingNormEquivAbsoluteValue.symm v x = v x := rfl + +end MulRingNorm_equiv_AbsoluteValue variable {R : Type*} [Ring R] /-- Two multiplicative ring norms `f, g` on `R` are equivalent if there exists a positive constant `c` such that for all `x ∈ R`, `(f x)^c = g x`. -/ - +@[deprecated "Use AbsoluteValue.IsEquiv instead" (since := "2025-01-07")] def equiv (f : MulRingNorm R) (g : MulRingNorm R) := ∃ c : ℝ, 0 < c ∧ (fun x => (f x) ^ c) = g +set_option linter.deprecated false in /-- Equivalence of multiplicative ring norms is reflexive. -/ +@[deprecated "Use AbsoluteValue.isEquiv_refl instead" (since := "2025-01-07")] lemma equiv_refl (f : MulRingNorm R) : equiv f f := by exact ⟨1, Real.zero_lt_one, by simp only [Real.rpow_one]⟩ +set_option linter.deprecated false in /-- Equivalence of multiplicative ring norms is symmetric. -/ +@[deprecated "Use AbsoluteValue.isEquiv_symm instead" (since := "2025-01-07")] lemma equiv_symm {f g : MulRingNorm R} (hfg : equiv f g) : equiv g f := by rcases hfg with ⟨c, hcpos, h⟩ use 1/c @@ -373,7 +414,9 @@ lemma equiv_symm {f g : MulRingNorm R} (hfg : equiv f g) : equiv g f := by ext x simpa [← congr_fun h x] using Real.rpow_rpow_inv (apply_nonneg f x) (ne_of_lt hcpos).symm +set_option linter.deprecated false in /-- Equivalence of multiplicative ring norms is transitive. -/ +@[deprecated "Use AbsoluteValue.isEquiv_trans instead" (since := "2025-01-07")] lemma equiv_trans {f g k : MulRingNorm R} (hfg : equiv f g) (hgk : equiv g k) : equiv f k := by rcases hfg with ⟨c, hcPos, hfg⟩ @@ -408,6 +451,7 @@ def normRingNorm (R : Type*) [NonUnitalNormedRing R] : RingNorm R := /-- A multiplicative ring norm satisfies `f n ≤ n` for every `n : ℕ`. -/ +@[deprecated "Use AbsoluteValue.apply_nat_le_self instead" (since := "2025-01-07")] lemma MulRingNorm_nat_le_nat {R : Type*} [Ring R] (n : ℕ) (f : MulRingNorm R) : f n ≤ n := by induction n with | zero => simp only [Nat.cast_zero, map_zero, le_refl] @@ -421,6 +465,7 @@ lemma MulRingNorm_nat_le_nat {R : Type*} [Ring R] (n : ℕ) (f : MulRingNorm R) open Int /-- A multiplicative norm composed with the absolute value on integers equals the norm itself. -/ +@[deprecated "Use AbsoluteValue.apply_natAbs_eq instead" (since := "2025-01-07")] lemma MulRingNorm.apply_natAbs_eq {R : Type*} [Ring R] (x : ℤ) (f : MulRingNorm R) : f (natAbs x) = f x := by obtain ⟨n, rfl | rfl⟩ := eq_nat_or_neg x <;> @@ -459,7 +504,16 @@ def NormedField.toMulRingNorm (R : Type*) [NormedField R] : MulRingNorm R where neg' := norm_neg eq_zero_of_map_eq_zero' x hx := by rw [← norm_eq_zero]; exact hx +/-- The norm on a `NormedField`, as an `AbsoluteValue`. -/ +def NormedField.toAbsoluteValue (R : Type*) [NormedField R] : AbsoluteValue R ℝ where + toFun := norm + map_mul' := norm_mul + nonneg' := norm_nonneg + eq_zero' _ := norm_eq_zero + add_le' := norm_add_le + /-- Triangle inequality for `MulRingNorm` applied to a list. -/ +@[deprecated "Use AbsoluteValue.listSum_le instead" (since := "2025-01-07")] lemma mulRingNorm_sum_le_sum_mulRingNorm {R : Type*} [NonAssocRing R] (l : List R) (f : MulRingNorm R) : f l.sum ≤ (l.map f).sum := by induction l with diff --git a/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean b/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean index 4a71d1dff9424..c04f7b9a2dcf5 100644 --- a/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean +++ b/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean @@ -374,29 +374,20 @@ theorem isLeast_opNorm (f : ContinuousMultilinearMap 𝕜 E G) : exact isClosed_Ici.inter (isClosed_iInter fun m ↦ isClosed_le continuous_const (continuous_id.mul continuous_const)) -@[deprecated (since := "2024-02-02")] alias isLeast_op_norm := isLeast_opNorm - theorem opNorm_nonneg (f : ContinuousMultilinearMap 𝕜 E G) : 0 ≤ ‖f‖ := Real.sInf_nonneg fun _ ⟨hx, _⟩ => hx -@[deprecated (since := "2024-02-02")] alias op_norm_nonneg := opNorm_nonneg - /-- The fundamental property of the operator norm of a continuous multilinear map: `‖f m‖` is bounded by `‖f‖` times the product of the `‖m i‖`. -/ theorem le_opNorm (f : ContinuousMultilinearMap 𝕜 E G) (m : ∀ i, E i) : ‖f m‖ ≤ ‖f‖ * ∏ i, ‖m i‖ := f.isLeast_opNorm.1.2 m -@[deprecated (since := "2024-02-02")] alias le_op_norm := le_opNorm - theorem le_mul_prod_of_opNorm_le_of_le {f : ContinuousMultilinearMap 𝕜 E G} {m : ∀ i, E i} {C : ℝ} {b : ι → ℝ} (hC : ‖f‖ ≤ C) (hm : ∀ i, ‖m i‖ ≤ b i) : ‖f m‖ ≤ C * ∏ i, b i := (f.le_opNorm m).trans <| by gcongr; exacts [f.opNorm_nonneg.trans hC, hm _] -@[deprecated (since := "2024-02-02")] -alias le_mul_prod_of_le_op_norm_of_le := le_mul_prod_of_opNorm_le_of_le - @[deprecated (since := "2024-11-27")] alias le_mul_prod_of_le_opNorm_of_le := le_mul_prod_of_opNorm_le_of_le @@ -404,67 +395,46 @@ theorem le_opNorm_mul_prod_of_le (f : ContinuousMultilinearMap 𝕜 E G) {m : ∀ i, E i} {b : ι → ℝ} (hm : ∀ i, ‖m i‖ ≤ b i) : ‖f m‖ ≤ ‖f‖ * ∏ i, b i := le_mul_prod_of_opNorm_le_of_le le_rfl hm -@[deprecated (since := "2024-02-02")] alias le_op_norm_mul_prod_of_le := le_opNorm_mul_prod_of_le - theorem le_opNorm_mul_pow_card_of_le (f : ContinuousMultilinearMap 𝕜 E G) {m b} (hm : ‖m‖ ≤ b) : ‖f m‖ ≤ ‖f‖ * b ^ Fintype.card ι := by simpa only [prod_const] using f.le_opNorm_mul_prod_of_le fun i => (norm_le_pi_norm m i).trans hm -@[deprecated (since := "2024-02-02")] -alias le_op_norm_mul_pow_card_of_le := le_opNorm_mul_pow_card_of_le - theorem le_opNorm_mul_pow_of_le {n : ℕ} {Ei : Fin n → Type*} [∀ i, SeminormedAddCommGroup (Ei i)] [∀ i, NormedSpace 𝕜 (Ei i)] (f : ContinuousMultilinearMap 𝕜 Ei G) {m : ∀ i, Ei i} {b : ℝ} (hm : ‖m‖ ≤ b) : ‖f m‖ ≤ ‖f‖ * b ^ n := by simpa only [Fintype.card_fin] using f.le_opNorm_mul_pow_card_of_le hm -@[deprecated (since := "2024-02-02")] alias le_op_norm_mul_pow_of_le := le_opNorm_mul_pow_of_le - theorem le_of_opNorm_le {f : ContinuousMultilinearMap 𝕜 E G} {C : ℝ} (h : ‖f‖ ≤ C) (m : ∀ i, E i) : ‖f m‖ ≤ C * ∏ i, ‖m i‖ := le_mul_prod_of_opNorm_le_of_le h fun _ ↦ le_rfl -@[deprecated (since := "2024-02-02")] alias le_of_op_norm_le := le_of_opNorm_le - theorem ratio_le_opNorm (f : ContinuousMultilinearMap 𝕜 E G) (m : ∀ i, E i) : (‖f m‖ / ∏ i, ‖m i‖) ≤ ‖f‖ := div_le_of_le_mul₀ (by positivity) (opNorm_nonneg _) (f.le_opNorm m) -@[deprecated (since := "2024-02-02")] alias ratio_le_op_norm := ratio_le_opNorm - /-- The image of the unit ball under a continuous multilinear map is bounded. -/ theorem unit_le_opNorm (f : ContinuousMultilinearMap 𝕜 E G) {m : ∀ i, E i} (h : ‖m‖ ≤ 1) : ‖f m‖ ≤ ‖f‖ := (le_opNorm_mul_pow_card_of_le f h).trans <| by simp -@[deprecated (since := "2024-02-02")] alias unit_le_op_norm := unit_le_opNorm - /-- If one controls the norm of every `f x`, then one controls the norm of `f`. -/ theorem opNorm_le_bound {f : ContinuousMultilinearMap 𝕜 E G} {M : ℝ} (hMp : 0 ≤ M) (hM : ∀ m, ‖f m‖ ≤ M * ∏ i, ‖m i‖) : ‖f‖ ≤ M := csInf_le bounds_bddBelow ⟨hMp, hM⟩ -@[deprecated (since := "2024-02-02")] alias op_norm_le_bound := opNorm_le_bound - theorem opNorm_le_iff {f : ContinuousMultilinearMap 𝕜 E G} {C : ℝ} (hC : 0 ≤ C) : ‖f‖ ≤ C ↔ ∀ m, ‖f m‖ ≤ C * ∏ i, ‖m i‖ := ⟨fun h _ ↦ le_of_opNorm_le h _, opNorm_le_bound hC⟩ -@[deprecated (since := "2024-02-02")] alias op_norm_le_iff := opNorm_le_iff - /-- The operator norm satisfies the triangle inequality. -/ theorem opNorm_add_le (f g : ContinuousMultilinearMap 𝕜 E G) : ‖f + g‖ ≤ ‖f‖ + ‖g‖ := opNorm_le_bound (add_nonneg (opNorm_nonneg f) (opNorm_nonneg g)) fun x => by rw [add_mul] exact norm_add_le_of_le (le_opNorm _ _) (le_opNorm _ _) -@[deprecated (since := "2024-02-02")] alias op_norm_add_le := opNorm_add_le - theorem opNorm_zero : ‖(0 : ContinuousMultilinearMap 𝕜 E G)‖ = 0 := (opNorm_nonneg _).antisymm' <| opNorm_le_bound le_rfl fun m => by simp -@[deprecated (since := "2024-02-02")] alias op_norm_zero := opNorm_zero - section variable {𝕜' : Type*} [NormedField 𝕜'] [NormedSpace 𝕜' G] [SMulCommClass 𝕜 𝕜' G] @@ -474,8 +444,6 @@ theorem opNorm_smul_le (c : 𝕜') (f : ContinuousMultilinearMap 𝕜 E G) : ‖ rw [smul_apply, norm_smul, mul_assoc] exact mul_le_mul_of_nonneg_left (le_opNorm _ _) (norm_nonneg _) -@[deprecated (since := "2024-02-02")] alias op_norm_smul_le := opNorm_smul_le - variable (𝕜 E G) in /-- Operator seminorm on the space of continuous multilinear maps, as `Seminorm`. @@ -544,8 +512,6 @@ instance normedSpace' : NormedSpace 𝕜' (ContinuousMultilinearMap 𝕜 (fun _ @[deprecated norm_neg (since := "2024-11-24")] theorem opNorm_neg (f : ContinuousMultilinearMap 𝕜 E G) : ‖-f‖ = ‖f‖ := norm_neg f -@[deprecated (since := "2024-02-02")] alias op_norm_neg := norm_neg - /-- The fundamental property of the operator norm of a continuous multilinear map: `‖f m‖` is bounded by `‖f‖` times the product of the `‖m i‖`, `nnnorm` version. -/ theorem le_opNNNorm (f : ContinuousMultilinearMap 𝕜 E G) (m : ∀ i, E i) : @@ -554,39 +520,27 @@ theorem le_opNNNorm (f : ContinuousMultilinearMap 𝕜 E G) (m : ∀ i, E i) : push_cast exact f.le_opNorm m -@[deprecated (since := "2024-02-02")] alias le_op_nnnorm := le_opNNNorm - theorem le_of_opNNNorm_le (f : ContinuousMultilinearMap 𝕜 E G) {C : ℝ≥0} (h : ‖f‖₊ ≤ C) (m : ∀ i, E i) : ‖f m‖₊ ≤ C * ∏ i, ‖m i‖₊ := (f.le_opNNNorm m).trans <| mul_le_mul' h le_rfl -@[deprecated (since := "2024-02-02")] alias le_of_op_nnnorm_le := le_of_opNNNorm_le - theorem opNNNorm_le_iff {f : ContinuousMultilinearMap 𝕜 E G} {C : ℝ≥0} : ‖f‖₊ ≤ C ↔ ∀ m, ‖f m‖₊ ≤ C * ∏ i, ‖m i‖₊ := by simp only [← NNReal.coe_le_coe]; simp [opNorm_le_iff C.coe_nonneg, NNReal.coe_prod] -@[deprecated (since := "2024-02-02")] alias op_nnnorm_le_iff := opNNNorm_le_iff - theorem isLeast_opNNNorm (f : ContinuousMultilinearMap 𝕜 E G) : IsLeast {C : ℝ≥0 | ∀ m, ‖f m‖₊ ≤ C * ∏ i, ‖m i‖₊} ‖f‖₊ := by simpa only [← opNNNorm_le_iff] using isLeast_Ici -@[deprecated (since := "2024-02-02")] alias isLeast_op_nnnorm := isLeast_opNNNorm - theorem opNNNorm_prod (f : ContinuousMultilinearMap 𝕜 E G) (g : ContinuousMultilinearMap 𝕜 E G') : ‖f.prod g‖₊ = max ‖f‖₊ ‖g‖₊ := eq_of_forall_ge_iff fun _ ↦ by - simp only [opNNNorm_le_iff, prod_apply, Prod.nnnorm_def', max_le_iff, forall_and] - -@[deprecated (since := "2024-02-02")] alias op_nnnorm_prod := opNNNorm_prod + simp only [opNNNorm_le_iff, prod_apply, Prod.nnnorm_def, max_le_iff, forall_and] theorem opNorm_prod (f : ContinuousMultilinearMap 𝕜 E G) (g : ContinuousMultilinearMap 𝕜 E G') : ‖f.prod g‖ = max ‖f‖ ‖g‖ := congr_arg NNReal.toReal (opNNNorm_prod f g) -@[deprecated (since := "2024-02-02")] alias op_norm_prod := opNorm_prod - theorem opNNNorm_pi [∀ i', SeminormedAddCommGroup (E' i')] [∀ i', NormedSpace 𝕜 (E' i')] (f : ∀ i', ContinuousMultilinearMap 𝕜 E (E' i')) : ‖pi f‖₊ = ‖f‖₊ := @@ -598,8 +552,6 @@ theorem opNorm_pi {ι' : Type v'} [Fintype ι'] {E' : ι' → Type wE'} ‖pi f‖ = ‖f‖ := congr_arg NNReal.toReal (opNNNorm_pi f) -@[deprecated (since := "2024-02-02")] alias op_norm_pi := opNorm_pi - section @[simp] @@ -658,14 +610,9 @@ variable (𝕜 E E' G G') def prodL : ContinuousMultilinearMap 𝕜 E G × ContinuousMultilinearMap 𝕜 E G' ≃ₗᵢ[𝕜] ContinuousMultilinearMap 𝕜 E (G × G') where - toFun f := f.1.prod f.2 - invFun f := - ((ContinuousLinearMap.fst 𝕜 G G').compContinuousMultilinearMap f, - (ContinuousLinearMap.snd 𝕜 G G').compContinuousMultilinearMap f) + __ := prodEquiv map_add' _ _ := rfl map_smul' _ _ := rfl - left_inv f := by ext <;> rfl - right_inv f := by ext <;> rfl norm_map' f := opNorm_prod f.1 f.2 /-- `ContinuousMultilinearMap.pi` as a `LinearIsometryEquiv`. -/ @@ -745,7 +692,7 @@ namespace ContinuousMultilinearMap these variables, and fixing the other ones equal to a given value `z`. It is denoted by `f.restr s hk z`, where `hk` is a proof that the cardinality of `s` is `k`. The implicit identification between `Fin k` and `s` that we use is the canonical (increasing) bijection. -/ -def restr {k n : ℕ} (f : (G[×n]→L[𝕜] G' : _)) (s : Finset (Fin n)) (hk : #s = k) (z : G) : +def restr {k n : ℕ} (f : (G[×n]→L[𝕜] G' :)) (s : Finset (Fin n)) (hk : #s = k) (z : G) : G[×k]→L[𝕜] G' := (f.toMultilinearMap.restr s hk z).mkContinuous (‖f‖ * ‖z‖ ^ (n - k)) fun _ => MultilinearMap.restr_norm_le _ _ _ _ f.le_opNorm _ @@ -1332,8 +1279,6 @@ variable {𝕜 : Type u} {ι : Type v} {E : ι → Type wE} {G : Type wG} {G' : theorem opNorm_zero_iff {f : ContinuousMultilinearMap 𝕜 E G} : ‖f‖ = 0 ↔ f = 0 := by simp [← (opNorm_nonneg f).le_iff_eq, opNorm_le_iff le_rfl, ContinuousMultilinearMap.ext_iff] -@[deprecated (since := "2024-02-02")] alias op_norm_zero_iff := opNorm_zero_iff - /-- Continuous multilinear maps themselves form a normed group with respect to the operator norm. -/ instance normedAddCommGroup : NormedAddCommGroup (ContinuousMultilinearMap 𝕜 E G) := diff --git a/Mathlib/Analysis/NormedSpace/OperatorNorm/Basic.lean b/Mathlib/Analysis/NormedSpace/OperatorNorm/Basic.lean index 34ac1d9ef91a5..bef49e1cf0329 100644 --- a/Mathlib/Analysis/NormedSpace/OperatorNorm/Basic.lean +++ b/Mathlib/Analysis/NormedSpace/OperatorNorm/Basic.lean @@ -136,14 +136,12 @@ theorem isLeast_opNorm [RingHomIsometric σ₁₂] (f : E →SL[σ₁₂] F) : simp only [setOf_and, setOf_forall] refine isClosed_Ici.inter <| isClosed_iInter fun _ ↦ isClosed_le ?_ ?_ <;> continuity -@[deprecated (since := "2024-02-02")] alias isLeast_op_norm := isLeast_opNorm /-- If one controls the norm of every `A x`, then one controls the norm of `A`. -/ theorem opNorm_le_bound (f : E →SL[σ₁₂] F) {M : ℝ} (hMp : 0 ≤ M) (hM : ∀ x, ‖f x‖ ≤ M * ‖x‖) : ‖f‖ ≤ M := csInf_le bounds_bddBelow ⟨hMp, hM⟩ -@[deprecated (since := "2024-02-02")] alias op_norm_le_bound := opNorm_le_bound /-- If one controls the norm of every `A x`, `‖x‖ ≠ 0`, then one controls the norm of `A`. -/ theorem opNorm_le_bound' (f : E →SL[σ₁₂] F) {M : ℝ} (hMp : 0 ≤ M) @@ -152,13 +150,11 @@ theorem opNorm_le_bound' (f : E →SL[σ₁₂] F) {M : ℝ} (hMp : 0 ≤ M) (ne_or_eq ‖x‖ 0).elim (hM x) fun h => by simp only [h, mul_zero, norm_image_of_norm_zero f f.2 h, le_refl] -@[deprecated (since := "2024-02-02")] alias op_norm_le_bound' := opNorm_le_bound' theorem opNorm_le_of_lipschitz {f : E →SL[σ₁₂] F} {K : ℝ≥0} (hf : LipschitzWith K f) : ‖f‖ ≤ K := f.opNorm_le_bound K.2 fun x => by simpa only [dist_zero_right, f.map_zero] using hf.dist_le_mul x 0 -@[deprecated (since := "2024-02-02")] alias op_norm_le_of_lipschitz := opNorm_le_of_lipschitz theorem opNorm_eq_of_bounds {φ : E →SL[σ₁₂] F} {M : ℝ} (M_nonneg : 0 ≤ M) (h_above : ∀ x, ‖φ x‖ ≤ M * ‖x‖) (h_below : ∀ N ≥ 0, (∀ x, ‖φ x‖ ≤ N * ‖x‖) → M ≤ N) : @@ -167,22 +163,18 @@ theorem opNorm_eq_of_bounds {φ : E →SL[σ₁₂] F} {M : ℝ} (M_nonneg : 0 ((le_csInf_iff ContinuousLinearMap.bounds_bddBelow ⟨M, M_nonneg, h_above⟩).mpr fun N ⟨N_nonneg, hN⟩ => h_below N N_nonneg hN) -@[deprecated (since := "2024-02-02")] alias op_norm_eq_of_bounds := opNorm_eq_of_bounds theorem opNorm_neg (f : E →SL[σ₁₂] F) : ‖-f‖ = ‖f‖ := by simp only [norm_def, neg_apply, norm_neg] -@[deprecated (since := "2024-02-02")] alias op_norm_neg := opNorm_neg theorem opNorm_nonneg (f : E →SL[σ₁₂] F) : 0 ≤ ‖f‖ := Real.sInf_nonneg fun _ ↦ And.left -@[deprecated (since := "2024-02-02")] alias op_norm_nonneg := opNorm_nonneg /-- The norm of the `0` operator is `0`. -/ theorem opNorm_zero : ‖(0 : E →SL[σ₁₂] F)‖ = 0 := le_antisymm (opNorm_le_bound _ le_rfl fun _ ↦ by simp) (opNorm_nonneg _) -@[deprecated (since := "2024-02-02")] alias op_norm_zero := opNorm_zero /-- The norm of the identity is at most `1`. It is in fact `1`, except when the space is trivial where it is `0`. It means that one can not do better than an inequality in general. -/ @@ -197,51 +189,42 @@ variable [RingHomIsometric σ₁₂] [RingHomIsometric σ₂₃] (f g : E →SL[ /-- The fundamental property of the operator norm: `‖f x‖ ≤ ‖f‖ * ‖x‖`. -/ theorem le_opNorm : ‖f x‖ ≤ ‖f‖ * ‖x‖ := (isLeast_opNorm f).1.2 x -@[deprecated (since := "2024-02-02")] alias le_op_norm := le_opNorm theorem dist_le_opNorm (x y : E) : dist (f x) (f y) ≤ ‖f‖ * dist x y := by simp_rw [dist_eq_norm, ← map_sub, f.le_opNorm] -@[deprecated (since := "2024-02-02")] alias dist_le_op_norm := dist_le_opNorm theorem le_of_opNorm_le_of_le {x} {a b : ℝ} (hf : ‖f‖ ≤ a) (hx : ‖x‖ ≤ b) : ‖f x‖ ≤ a * b := (f.le_opNorm x).trans <| by gcongr; exact (opNorm_nonneg f).trans hf -@[deprecated (since := "2024-02-02")] alias le_of_op_norm_le_of_le := le_of_opNorm_le_of_le theorem le_opNorm_of_le {c : ℝ} {x} (h : ‖x‖ ≤ c) : ‖f x‖ ≤ ‖f‖ * c := f.le_of_opNorm_le_of_le le_rfl h -@[deprecated (since := "2024-02-02")] alias le_op_norm_of_le := le_opNorm_of_le theorem le_of_opNorm_le {c : ℝ} (h : ‖f‖ ≤ c) (x : E) : ‖f x‖ ≤ c * ‖x‖ := f.le_of_opNorm_le_of_le h le_rfl -@[deprecated (since := "2024-02-02")] alias le_of_op_norm_le := le_of_opNorm_le theorem opNorm_le_iff {f : E →SL[σ₁₂] F} {M : ℝ} (hMp : 0 ≤ M) : ‖f‖ ≤ M ↔ ∀ x, ‖f x‖ ≤ M * ‖x‖ := ⟨f.le_of_opNorm_le, opNorm_le_bound f hMp⟩ -@[deprecated (since := "2024-02-02")] alias op_norm_le_iff := opNorm_le_iff theorem ratio_le_opNorm : ‖f x‖ / ‖x‖ ≤ ‖f‖ := div_le_of_le_mul₀ (norm_nonneg _) f.opNorm_nonneg (le_opNorm _ _) -@[deprecated (since := "2024-02-02")] alias ratio_le_op_norm := ratio_le_opNorm /-- The image of the unit ball under a continuous linear map is bounded. -/ theorem unit_le_opNorm : ‖x‖ ≤ 1 → ‖f x‖ ≤ ‖f‖ := mul_one ‖f‖ ▸ f.le_opNorm_of_le -@[deprecated (since := "2024-02-02")] alias unit_le_op_norm := unit_le_opNorm theorem opNorm_le_of_shell {f : E →SL[σ₁₂] F} {ε C : ℝ} (ε_pos : 0 < ε) (hC : 0 ≤ C) {c : 𝕜} (hc : 1 < ‖c‖) (hf : ∀ x, ε / ‖c‖ ≤ ‖x‖ → ‖x‖ < ε → ‖f x‖ ≤ C * ‖x‖) : ‖f‖ ≤ C := f.opNorm_le_bound' hC fun _ hx => SemilinearMapClass.bound_of_shell_semi_normed f ε_pos hc hf hx -@[deprecated (since := "2024-02-02")] alias op_norm_le_of_shell := opNorm_le_of_shell theorem opNorm_le_of_ball {f : E →SL[σ₁₂] F} {ε : ℝ} {C : ℝ} (ε_pos : 0 < ε) (hC : 0 ≤ C) (hf : ∀ x ∈ ball (0 : E) ε, ‖f x‖ ≤ C * ‖x‖) : ‖f‖ ≤ C := by @@ -249,14 +232,12 @@ theorem opNorm_le_of_ball {f : E →SL[σ₁₂] F} {ε : ℝ} {C : ℝ} (ε_pos refine opNorm_le_of_shell ε_pos hC hc fun x _ hx => hf x ?_ rwa [ball_zero_eq] -@[deprecated (since := "2024-02-02")] alias op_norm_le_of_ball := opNorm_le_of_ball theorem opNorm_le_of_nhds_zero {f : E →SL[σ₁₂] F} {C : ℝ} (hC : 0 ≤ C) (hf : ∀ᶠ x in 𝓝 (0 : E), ‖f x‖ ≤ C * ‖x‖) : ‖f‖ ≤ C := let ⟨_, ε0, hε⟩ := Metric.eventually_nhds_iff_ball.1 hf opNorm_le_of_ball ε0 hC hε -@[deprecated (since := "2024-02-02")] alias op_norm_le_of_nhds_zero := opNorm_le_of_nhds_zero theorem opNorm_le_of_shell' {f : E →SL[σ₁₂] F} {ε C : ℝ} (ε_pos : 0 < ε) (hC : 0 ≤ C) {c : 𝕜} (hc : ‖c‖ < 1) (hf : ∀ x, ε * ‖c‖ ≤ ‖x‖ → ‖x‖ < ε → ‖f x‖ ≤ C * ‖x‖) : ‖f‖ ≤ C := by @@ -268,7 +249,6 @@ theorem opNorm_le_of_shell' {f : E →SL[σ₁₂] F} {ε C : ℝ} (ε_pos : 0 < refine opNorm_le_of_shell ε_pos hC hc ?_ rwa [norm_inv, div_eq_mul_inv, inv_inv] -@[deprecated (since := "2024-02-02")] alias op_norm_le_of_shell' := opNorm_le_of_shell' /-- For a continuous real linear map `f`, if one controls the norm of every `f x`, `‖x‖ = 1`, then one controls the norm of `f`. -/ @@ -280,14 +260,12 @@ theorem opNorm_le_of_unit_norm [NormedSpace ℝ E] [NormedSpace ℝ F] {f : E rwa [map_smul, norm_smul, norm_inv, norm_norm, ← div_eq_inv_mul, div_le_iff₀] at H₂ exact (norm_nonneg x).lt_of_ne' hx -@[deprecated (since := "2024-02-02")] alias op_norm_le_of_unit_norm := opNorm_le_of_unit_norm /-- The operator norm satisfies the triangle inequality. -/ theorem opNorm_add_le : ‖f + g‖ ≤ ‖f‖ + ‖g‖ := (f + g).opNorm_le_bound (add_nonneg f.opNorm_nonneg g.opNorm_nonneg) fun x => (norm_add_le_of_le (f.le_opNorm x) (g.le_opNorm x)).trans_eq (add_mul _ _ _).symm -@[deprecated (since := "2024-02-02")] alias op_norm_add_le := opNorm_add_le /-- If there is an element with norm different from `0`, then the norm of the identity equals `1`. (Since we are working with seminorms supposing that the space is non-trivial is not enough.) -/ @@ -303,7 +281,6 @@ theorem opNorm_smul_le {𝕜' : Type*} [NormedField 𝕜'] [NormedSpace 𝕜' F] rw [smul_apply, norm_smul, mul_assoc] exact mul_le_mul_of_nonneg_left (le_opNorm _ _) (norm_nonneg _) -@[deprecated (since := "2024-02-02")] alias op_norm_smul_le := opNorm_smul_le /-- Operator seminorm on the space of continuous (semi)linear maps, as `Seminorm`. @@ -352,7 +329,6 @@ theorem opNorm_comp_le (f : E →SL[σ₁₂] F) : ‖h.comp f‖ ≤ ‖h‖ * rw [mul_assoc] exact h.le_opNorm_of_le (f.le_opNorm x)⟩ -@[deprecated (since := "2024-02-02")] alias op_norm_comp_le := opNorm_comp_le /-- Continuous linear maps form a seminormed ring with respect to the operator norm. -/ instance toSemiNormedRing : SeminormedRing (E →L[𝕜] E) := @@ -378,7 +354,6 @@ theorem opNorm_subsingleton [Subsingleton E] : ‖f‖ = 0 := by intro x simp [Subsingleton.elim x 0] -@[deprecated (since := "2024-02-02")] alias op_norm_subsingleton := opNorm_subsingleton end OpNorm diff --git a/Mathlib/Analysis/NormedSpace/OperatorNorm/Bilinear.lean b/Mathlib/Analysis/NormedSpace/OperatorNorm/Bilinear.lean index 9cb3ac62feea8..8193941231cfa 100644 --- a/Mathlib/Analysis/NormedSpace/OperatorNorm/Bilinear.lean +++ b/Mathlib/Analysis/NormedSpace/OperatorNorm/Bilinear.lean @@ -55,7 +55,6 @@ theorem opNorm_ext [RingHomIsometric σ₁₃] (f : E →SL[σ₁₂] F) (g : E rw [← h z] exact h₂ z -@[deprecated (since := "2024-02-02")] alias op_norm_ext := opNorm_ext variable [RingHomIsometric σ₂₃] @@ -63,20 +62,17 @@ theorem opNorm_le_bound₂ (f : E →SL[σ₁₃] F →SL[σ₂₃] G) {C : ℝ} (hC : ∀ x y, ‖f x y‖ ≤ C * ‖x‖ * ‖y‖) : ‖f‖ ≤ C := f.opNorm_le_bound h0 fun x => (f x).opNorm_le_bound (mul_nonneg h0 (norm_nonneg _)) <| hC x -@[deprecated (since := "2024-02-02")] alias op_norm_le_bound₂ := opNorm_le_bound₂ theorem le_opNorm₂ [RingHomIsometric σ₁₃] (f : E →SL[σ₁₃] F →SL[σ₂₃] G) (x : E) (y : F) : ‖f x y‖ ≤ ‖f‖ * ‖x‖ * ‖y‖ := (f x).le_of_opNorm_le (f.le_opNorm x) y -@[deprecated (since := "2024-02-02")] alias le_op_norm₂ := le_opNorm₂ theorem le_of_opNorm₂_le_of_le [RingHomIsometric σ₁₃] (f : E →SL[σ₁₃] F →SL[σ₂₃] G) {x : E} {y : F} {a b c : ℝ} (hf : ‖f‖ ≤ a) (hx : ‖x‖ ≤ b) (hy : ‖y‖ ≤ c) : ‖f x y‖ ≤ a * b * c := (f x).le_of_opNorm_le_of_le (f.le_of_opNorm_le_of_le hf hx) hy -@[deprecated (since := "2024-02-02")] alias le_of_op_norm₂_le_of_le := le_of_opNorm₂_le_of_le end OpNorm @@ -172,7 +168,6 @@ theorem flip_flip (f : E →SL[σ₁₃] F →SL[σ₂₃] G) : f.flip.flip = f theorem opNorm_flip (f : E →SL[σ₁₃] F →SL[σ₂₃] G) : ‖f.flip‖ = ‖f‖ := le_antisymm (by simpa only [flip_flip] using le_norm_flip f.flip) (le_norm_flip f) -@[deprecated (since := "2024-02-02")] alias op_norm_flip := opNorm_flip @[simp] theorem flip_add (f g : E →SL[σ₁₃] F →SL[σ₂₃] G) : (f + g).flip = f.flip + g.flip := diff --git a/Mathlib/Analysis/NormedSpace/OperatorNorm/Completeness.lean b/Mathlib/Analysis/NormedSpace/OperatorNorm/Completeness.lean index 4043b6833c43d..39bbe337bb7d7 100644 --- a/Mathlib/Analysis/NormedSpace/OperatorNorm/Completeness.lean +++ b/Mathlib/Analysis/NormedSpace/OperatorNorm/Completeness.lean @@ -247,7 +247,6 @@ theorem opNorm_extend_le : _ ≤ ‖f‖ * (N * ‖e x‖) := mul_le_mul_of_nonneg_left (h_e x) (norm_nonneg _) _ ≤ N * ‖f‖ * ‖e x‖ := by rw [mul_comm ↑N ‖f‖, mul_assoc] -@[deprecated (since := "2024-02-02")] alias op_norm_extend_le := opNorm_extend_le end diff --git a/Mathlib/Analysis/NormedSpace/OperatorNorm/Mul.lean b/Mathlib/Analysis/NormedSpace/OperatorNorm/Mul.lean index c45a659de2c23..147350c036a57 100644 --- a/Mathlib/Analysis/NormedSpace/OperatorNorm/Mul.lean +++ b/Mathlib/Analysis/NormedSpace/OperatorNorm/Mul.lean @@ -45,12 +45,10 @@ theorem mul_apply' (x y : 𝕜') : mul 𝕜 𝕜' x y = x * y := theorem opNorm_mul_apply_le (x : 𝕜') : ‖mul 𝕜 𝕜' x‖ ≤ ‖x‖ := opNorm_le_bound _ (norm_nonneg x) (norm_mul_le x) -@[deprecated (since := "2024-02-02")] alias op_norm_mul_apply_le := opNorm_mul_apply_le theorem opNorm_mul_le : ‖mul 𝕜 𝕜'‖ ≤ 1 := LinearMap.mkContinuous₂_norm_le _ zero_le_one _ -@[deprecated (since := "2024-02-02")] alias op_norm_mul_le := opNorm_mul_le /-- Multiplication on the left in a non-unital normed algebra `𝕜'` as a non-unital algebra homomorphism into the algebra of *continuous* linear maps. This is the left regular representation @@ -87,22 +85,14 @@ theorem opNorm_mulLeftRight_apply_apply_le (x y : 𝕜') : ‖mulLeftRight 𝕜 (opNorm_le_bound _ (norm_nonneg _) fun _ => (norm_mul_le _ _).trans_eq (mul_comm _ _)) (norm_nonneg _) (norm_nonneg _) -@[deprecated (since := "2024-02-02")] -alias op_norm_mulLeftRight_apply_apply_le := - opNorm_mulLeftRight_apply_apply_le - theorem opNorm_mulLeftRight_apply_le (x : 𝕜') : ‖mulLeftRight 𝕜 𝕜' x‖ ≤ ‖x‖ := opNorm_le_bound _ (norm_nonneg x) (opNorm_mulLeftRight_apply_apply_le 𝕜 𝕜' x) -@[deprecated (since := "2024-02-02")] -alias op_norm_mulLeftRight_apply_le := opNorm_mulLeftRight_apply_le - set_option maxSynthPendingDepth 2 in theorem opNorm_mulLeftRight_le : ‖mulLeftRight 𝕜 𝕜'‖ ≤ 1 := opNorm_le_bound _ zero_le_one fun x => (one_mul ‖x‖).symm ▸ opNorm_mulLeftRight_apply_le 𝕜 𝕜' x -@[deprecated (since := "2024-02-02")] alias op_norm_mulLeftRight_le := opNorm_mulLeftRight_le /-- This is a mixin class for non-unital normed algebras which states that the left-regular representation of the algebra on itself is isometric. Every unital normed algebra with `‖1‖ = 1` is @@ -133,13 +123,11 @@ lemma isometry_mul : Isometry (mul 𝕜 𝕜') := lemma opNorm_mul_apply (x : 𝕜') : ‖mul 𝕜 𝕜' x‖ = ‖x‖ := (AddMonoidHomClass.isometry_iff_norm (mul 𝕜 𝕜')).mp (isometry_mul 𝕜 𝕜') x -@[deprecated (since := "2024-02-02")] alias op_norm_mul_apply := opNorm_mul_apply @[simp] lemma opNNNorm_mul_apply (x : 𝕜') : ‖mul 𝕜 𝕜' x‖₊ = ‖x‖₊ := Subtype.ext <| opNorm_mul_apply 𝕜 𝕜' x -@[deprecated (since := "2024-02-02")] alias op_nnnorm_mul_apply := opNNNorm_mul_apply /-- Multiplication in a normed algebra as a linear isometry to the space of continuous linear maps. -/ @@ -210,7 +198,6 @@ variable {𝕜} theorem opNorm_lsmul_apply_le (x : 𝕜') : ‖(lsmul 𝕜 𝕜' x : E →L[𝕜] E)‖ ≤ ‖x‖ := ContinuousLinearMap.opNorm_le_bound _ (norm_nonneg x) fun y => norm_smul_le x y -@[deprecated (since := "2024-02-02")] alias op_norm_lsmul_apply_le := opNorm_lsmul_apply_le /-- The norm of `lsmul` is at most 1 in any semi-normed group. -/ theorem opNorm_lsmul_le : ‖(lsmul 𝕜 𝕜' : 𝕜' →L[𝕜] E →L[𝕜] E)‖ ≤ 1 := by @@ -218,7 +205,6 @@ theorem opNorm_lsmul_le : ‖(lsmul 𝕜 𝕜' : 𝕜' →L[𝕜] E →L[𝕜] E simp_rw [one_mul] exact opNorm_lsmul_apply_le _ -@[deprecated (since := "2024-02-02")] alias op_norm_lsmul_le := opNorm_lsmul_le end SMulLinear @@ -242,13 +228,11 @@ variable [SMulCommClass 𝕜 𝕜' 𝕜'] [RegularNormedAlgebra 𝕜 𝕜'] [Non theorem opNorm_mul : ‖mul 𝕜 𝕜'‖ = 1 := (mulₗᵢ 𝕜 𝕜').norm_toContinuousLinearMap -@[deprecated (since := "2024-02-02")] alias op_norm_mul := opNorm_mul @[simp] theorem opNNNorm_mul : ‖mul 𝕜 𝕜'‖₊ = 1 := Subtype.ext <| opNorm_mul 𝕜 𝕜' -@[deprecated (since := "2024-02-02")] alias op_nnnorm_mul := opNNNorm_mul end @@ -267,7 +251,6 @@ theorem opNorm_lsmul [NormedField 𝕜'] [NormedAlgebra 𝕜 𝕜'] [NormedSpace refine le_of_mul_le_mul_right ?_ (norm_pos_iff.mpr hy) simp_rw [one_mul, this] -@[deprecated (since := "2024-02-02")] alias op_norm_lsmul := opNorm_lsmul end ContinuousLinearMap diff --git a/Mathlib/Analysis/NormedSpace/OperatorNorm/NNNorm.lean b/Mathlib/Analysis/NormedSpace/OperatorNorm/NNNorm.lean index 51f3b1705bfc9..018b01d78582b 100644 --- a/Mathlib/Analysis/NormedSpace/OperatorNorm/NNNorm.lean +++ b/Mathlib/Analysis/NormedSpace/OperatorNorm/NNNorm.lean @@ -56,14 +56,12 @@ theorem nnnorm_def (f : E →SL[σ₁₂] F) : ‖f‖₊ = sInf { c | ∀ x, theorem opNNNorm_le_bound (f : E →SL[σ₁₂] F) (M : ℝ≥0) (hM : ∀ x, ‖f x‖₊ ≤ M * ‖x‖₊) : ‖f‖₊ ≤ M := opNorm_le_bound f (zero_le M) hM -@[deprecated (since := "2024-02-02")] alias op_nnnorm_le_bound := opNNNorm_le_bound /-- If one controls the norm of every `A x`, `‖x‖₊ ≠ 0`, then one controls the norm of `A`. -/ theorem opNNNorm_le_bound' (f : E →SL[σ₁₂] F) (M : ℝ≥0) (hM : ∀ x, ‖x‖₊ ≠ 0 → ‖f x‖₊ ≤ M * ‖x‖₊) : ‖f‖₊ ≤ M := opNorm_le_bound' f (zero_le M) fun x hx => hM x <| by rwa [← NNReal.coe_ne_zero] -@[deprecated (since := "2024-02-02")] alias op_nnnorm_le_bound' := opNNNorm_le_bound' /-- For a continuous real linear map `f`, if one controls the norm of every `f x`, `‖x‖₊ = 1`, then one controls the norm of `f`. -/ @@ -71,45 +69,35 @@ theorem opNNNorm_le_of_unit_nnnorm [NormedSpace ℝ E] [NormedSpace ℝ F] {f : (hf : ∀ x, ‖x‖₊ = 1 → ‖f x‖₊ ≤ C) : ‖f‖₊ ≤ C := opNorm_le_of_unit_norm C.coe_nonneg fun x hx => hf x <| by rwa [← NNReal.coe_eq_one] -@[deprecated (since := "2024-02-02")] -alias op_nnnorm_le_of_unit_nnnorm := opNNNorm_le_of_unit_nnnorm - theorem opNNNorm_le_of_lipschitz {f : E →SL[σ₁₂] F} {K : ℝ≥0} (hf : LipschitzWith K f) : ‖f‖₊ ≤ K := opNorm_le_of_lipschitz hf -@[deprecated (since := "2024-02-02")] alias op_nnnorm_le_of_lipschitz := opNNNorm_le_of_lipschitz theorem opNNNorm_eq_of_bounds {φ : E →SL[σ₁₂] F} (M : ℝ≥0) (h_above : ∀ x, ‖φ x‖₊ ≤ M * ‖x‖₊) (h_below : ∀ N, (∀ x, ‖φ x‖₊ ≤ N * ‖x‖₊) → M ≤ N) : ‖φ‖₊ = M := Subtype.ext <| opNorm_eq_of_bounds (zero_le M) h_above <| Subtype.forall'.mpr h_below -@[deprecated (since := "2024-02-02")] alias op_nnnorm_eq_of_bounds := opNNNorm_eq_of_bounds theorem opNNNorm_le_iff {f : E →SL[σ₁₂] F} {C : ℝ≥0} : ‖f‖₊ ≤ C ↔ ∀ x, ‖f x‖₊ ≤ C * ‖x‖₊ := opNorm_le_iff C.2 -@[deprecated (since := "2024-02-02")] alias op_nnnorm_le_iff := opNNNorm_le_iff theorem isLeast_opNNNorm : IsLeast {C : ℝ≥0 | ∀ x, ‖f x‖₊ ≤ C * ‖x‖₊} ‖f‖₊ := by simpa only [← opNNNorm_le_iff] using isLeast_Ici -@[deprecated (since := "2024-02-02")] alias isLeast_op_nnnorm := isLeast_opNNNorm theorem opNNNorm_comp_le [RingHomIsometric σ₁₃] (f : E →SL[σ₁₂] F) : ‖h.comp f‖₊ ≤ ‖h‖₊ * ‖f‖₊ := opNorm_comp_le h f -@[deprecated (since := "2024-02-02")] alias op_nnnorm_comp_le := opNNNorm_comp_le theorem le_opNNNorm : ‖f x‖₊ ≤ ‖f‖₊ * ‖x‖₊ := f.le_opNorm x -@[deprecated (since := "2024-02-02")] alias le_op_nnnorm := le_opNNNorm theorem nndist_le_opNNNorm (x y : E) : nndist (f x) (f y) ≤ ‖f‖₊ * nndist x y := dist_le_opNorm f x y -@[deprecated (since := "2024-02-02")] alias nndist_le_op_nnnorm := nndist_le_opNNNorm /-- continuous linear maps are Lipschitz continuous. -/ theorem lipschitz : LipschitzWith ‖f‖₊ f := @@ -131,17 +119,11 @@ theorem exists_mul_lt_apply_of_lt_opNNNorm (f : E →SL[σ₁₂] F) {r : ℝ≥ not_mem_of_lt_csInf (nnnorm_def f ▸ hr : r < sInf { c : ℝ≥0 | ∀ x, ‖f x‖₊ ≤ c * ‖x‖₊ }) (OrderBot.bddBelow _) -@[deprecated (since := "2024-02-02")] -alias exists_mul_lt_apply_of_lt_op_nnnorm := exists_mul_lt_apply_of_lt_opNNNorm - theorem exists_mul_lt_of_lt_opNorm (f : E →SL[σ₁₂] F) {r : ℝ} (hr₀ : 0 ≤ r) (hr : r < ‖f‖) : ∃ x, r * ‖x‖ < ‖f x‖ := by lift r to ℝ≥0 using hr₀ exact f.exists_mul_lt_apply_of_lt_opNNNorm hr -@[deprecated (since := "2024-02-02")] -alias exists_mul_lt_of_lt_op_norm := exists_mul_lt_of_lt_opNorm - theorem exists_lt_apply_of_lt_opNNNorm {𝕜 𝕜₂ E F : Type*} [NormedAddCommGroup E] [SeminormedAddCommGroup F] [DenselyNormedField 𝕜] [NontriviallyNormedField 𝕜₂] {σ₁₂ : 𝕜 →+* 𝕜₂} [NormedSpace 𝕜 E] [NormedSpace 𝕜₂ F] [RingHomIsometric σ₁₂] (f : E →SL[σ₁₂] F) {r : ℝ≥0} @@ -158,9 +140,6 @@ theorem exists_lt_apply_of_lt_opNNNorm {𝕜 𝕜₂ E F : Type*} [NormedAddComm have : ‖σ₁₂ k‖₊ = ‖k‖₊ := Subtype.ext RingHomIsometric.is_iso rwa [map_smulₛₗ f, nnnorm_smul, ← div_lt_iff₀ hfy.bot_lt, div_eq_mul_inv, this] -@[deprecated (since := "2024-02-02")] -alias exists_lt_apply_of_lt_op_nnnorm := exists_lt_apply_of_lt_opNNNorm - theorem exists_lt_apply_of_lt_opNorm {𝕜 𝕜₂ E F : Type*} [NormedAddCommGroup E] [SeminormedAddCommGroup F] [DenselyNormedField 𝕜] [NontriviallyNormedField 𝕜₂] {σ₁₂ : 𝕜 →+* 𝕜₂} [NormedSpace 𝕜 E] [NormedSpace 𝕜₂ F] [RingHomIsometric σ₁₂] (f : E →SL[σ₁₂] F) {r : ℝ} @@ -170,9 +149,6 @@ theorem exists_lt_apply_of_lt_opNorm {𝕜 𝕜₂ E F : Type*} [NormedAddCommGr · lift r to ℝ≥0 using not_lt.1 hr₀ exact f.exists_lt_apply_of_lt_opNNNorm hr -@[deprecated (since := "2024-02-02")] -alias exists_lt_apply_of_lt_op_norm := exists_lt_apply_of_lt_opNorm - theorem sSup_unit_ball_eq_nnnorm {𝕜 𝕜₂ E F : Type*} [NormedAddCommGroup E] [SeminormedAddCommGroup F] [DenselyNormedField 𝕜] [NontriviallyNormedField 𝕜₂] {σ₁₂ : 𝕜 →+* 𝕜₂} [NormedSpace 𝕜 E] [NormedSpace 𝕜₂ F] [RingHomIsometric σ₁₂] (f : E →SL[σ₁₂] F) : diff --git a/Mathlib/Analysis/NormedSpace/OperatorNorm/NormedSpace.lean b/Mathlib/Analysis/NormedSpace/OperatorNorm/NormedSpace.lean index b08ba685f3320..f9b408f94f338 100644 --- a/Mathlib/Analysis/NormedSpace/OperatorNorm/NormedSpace.lean +++ b/Mathlib/Analysis/NormedSpace/OperatorNorm/NormedSpace.lean @@ -102,7 +102,6 @@ theorem opNorm_zero_iff [RingHomIsometric σ₁₂] : ‖f‖ = 0 ↔ f = 0 := rintro rfl exact opNorm_zero) -@[deprecated (since := "2024-02-02")] alias op_norm_zero_iff := opNorm_zero_iff /-- If a normed space is non-trivial, then the norm of the identity equals `1`. -/ @[simp] @@ -208,9 +207,6 @@ theorem opNorm_comp_linearIsometryEquiv (f : F →SL[σ₂₃] G) (g : F' ≃ₛ haveI := g.symm.surjective.nontrivial simp [g.symm.toLinearIsometry.norm_toContinuousLinearMap] -@[deprecated (since := "2024-02-02")] -alias op_norm_comp_linearIsometryEquiv := opNorm_comp_linearIsometryEquiv - @[simp] theorem norm_smulRightL (c : E →L[𝕜] 𝕜) [Nontrivial Fₗ] : ‖smulRightL 𝕜 E Fₗ c‖ = ‖c‖ := ContinuousLinearMap.homothety_norm _ c.norm_smulRight_apply diff --git a/Mathlib/Analysis/NormedSpace/OperatorNorm/Prod.lean b/Mathlib/Analysis/NormedSpace/OperatorNorm/Prod.lean index 51e531175bf03..f50d7b58dac0e 100644 --- a/Mathlib/Analysis/NormedSpace/OperatorNorm/Prod.lean +++ b/Mathlib/Analysis/NormedSpace/OperatorNorm/Prod.lean @@ -52,13 +52,11 @@ theorem opNorm_prod (f : E →L[𝕜] F) (g : E →L[𝕜] G) : ‖f.prod g‖ = (opNorm_le_bound _ (norm_nonneg _) fun x => (le_max_right _ _).trans ((f.prod g).le_opNorm x)) -@[deprecated (since := "2024-02-02")] alias op_norm_prod := opNorm_prod @[simp] theorem opNNNorm_prod (f : E →L[𝕜] F) (g : E →L[𝕜] G) : ‖f.prod g‖₊ = ‖(f, g)‖₊ := Subtype.ext <| opNorm_prod f g -@[deprecated (since := "2024-02-02")] alias op_nnnorm_prod := opNNNorm_prod /-- `ContinuousLinearMap.prod` as a `LinearIsometryEquiv`. -/ def prodₗᵢ (R : Type*) [Semiring R] [Module R F] [Module R G] [ContinuousConstSMul R F] @@ -122,7 +120,7 @@ theorem _root_.Continuous.prod_map_equivL {f : X → M₁ ≃L[𝕜] M₂} {g : theorem _root_.ContinuousOn.prod_mapL {f : X → M₁ →L[𝕜] M₂} {g : X → M₃ →L[𝕜] M₄} {s : Set X} (hf : ContinuousOn f s) (hg : ContinuousOn g s) : ContinuousOn (fun x => (f x).prodMap (g x)) s := - ((prodMapL 𝕜 M₁ M₂ M₃ M₄).continuous.comp_continuousOn (hf.prod hg) : _) + ((prodMapL 𝕜 M₁ M₂ M₃ M₄).continuous.comp_continuousOn (hf.prod hg) :) theorem _root_.ContinuousOn.prod_map_equivL {f : X → M₁ ≃L[𝕜] M₂} {g : X → M₃ ≃L[𝕜] M₄} {s : Set X} (hf : ContinuousOn (fun x => (f x : M₁ →L[𝕜] M₂)) s) diff --git a/Mathlib/Analysis/NormedSpace/RCLike.lean b/Mathlib/Analysis/NormedSpace/RCLike.lean index f578908f625cc..a918890d4c64b 100644 --- a/Mathlib/Analysis/NormedSpace/RCLike.lean +++ b/Mathlib/Analysis/NormedSpace/RCLike.lean @@ -85,10 +85,6 @@ theorem ContinuousLinearMap.opNorm_bound_of_ball_bound {r : ℝ} (r_pos : 0 < r) apply LinearMap.bound_of_ball_bound' r_pos exact fun z hz => h z hz -@[deprecated (since := "2024-02-02")] -alias ContinuousLinearMap.op_norm_bound_of_ball_bound := - ContinuousLinearMap.opNorm_bound_of_ball_bound - variable (𝕜) include 𝕜 in theorem NormedSpace.sphere_nonempty_rclike [Nontrivial E] {r : ℝ} (hr : 0 ≤ r) : diff --git a/Mathlib/Analysis/SpecialFunctions/Complex/Arg.lean b/Mathlib/Analysis/SpecialFunctions/Complex/Arg.lean index 0086e643a19b8..1b221e58e8826 100644 --- a/Mathlib/Analysis/SpecialFunctions/Complex/Arg.lean +++ b/Mathlib/Analysis/SpecialFunctions/Complex/Arg.lean @@ -201,7 +201,7 @@ lemma natCast_arg {n : ℕ} : arg n = 0 := ofReal_natCast n ▸ arg_ofReal_of_nonneg n.cast_nonneg @[simp] -lemma ofNat_arg {n : ℕ} [n.AtLeastTwo] : arg (no_index (OfNat.ofNat n)) = 0 := +lemma ofNat_arg {n : ℕ} [n.AtLeastTwo] : arg ofNat(n) = 0 := natCast_arg theorem arg_eq_zero_iff {z : ℂ} : arg z = 0 ↔ 0 ≤ z.re ∧ z.im = 0 := by diff --git a/Mathlib/Analysis/SpecialFunctions/Complex/Log.lean b/Mathlib/Analysis/SpecialFunctions/Complex/Log.lean index 6beddd4acf3e0..60934d34bd710 100644 --- a/Mathlib/Analysis/SpecialFunctions/Complex/Log.lean +++ b/Mathlib/Analysis/SpecialFunctions/Complex/Log.lean @@ -65,7 +65,7 @@ lemma natCast_log {n : ℕ} : Real.log n = log n := ofReal_natCast n ▸ ofReal_ @[simp] lemma ofNat_log {n : ℕ} [n.AtLeastTwo] : - Real.log (no_index (OfNat.ofNat n)) = log (OfNat.ofNat n) := + Real.log ofNat(n) = log (OfNat.ofNat n) := natCast_log theorem log_ofReal_re (x : ℝ) : (log (x : ℂ)).re = Real.log x := by simp [log_re] diff --git a/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/ExpLog.lean b/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/ExpLog.lean index 27a9142d793a9..bdb5875943ebe 100644 --- a/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/ExpLog.lean +++ b/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/ExpLog.lean @@ -152,7 +152,7 @@ lemma log_pow (n : ℕ) (a : A) (ha₂ : ∀ x ∈ spectrum ℝ a, 0 < x) variable [CompleteSpace A] lemma log_exp (a : A) (ha : IsSelfAdjoint a := by cfc_tac) : log (NormedSpace.exp ℝ a) = a := by - have hcont : ContinuousOn Real.log (Real.exp '' spectrum ℝ a) := by fun_prop (disch := aesop) + have hcont : ContinuousOn Real.log (Real.exp '' spectrum ℝ a) := by fun_prop (disch := simp) rw [log, ← real_exp_eq_normedSpace_exp, ← cfc_comp' Real.log Real.exp a hcont] simp [cfc_id' (R := ℝ) a] diff --git a/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/Rpow.lean b/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/Rpow.lean index 487d14cd2871b..0bdd483e3237a 100644 --- a/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/Rpow.lean +++ b/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/Rpow.lean @@ -365,7 +365,7 @@ lemma sqrt_rpow_nnreal {a : A} {x : ℝ≥0} : sqrt (a ^ (x : ℝ)) = a ^ (x / 2 by_cases hx : x = 0 case pos => simp [hx, rpow_zero _ htriv] case neg => - have h₁ : 0 < x := lt_of_le_of_ne (by aesop) (Ne.symm hx) + have h₁ : 0 < x := lt_of_le_of_ne (by simp) (Ne.symm hx) have h₂ : (x : ℝ) / 2 = NNReal.toReal (x / 2) := rfl have h₃ : 0 < x / 2 := by positivity rw [← nnrpow_eq_rpow h₁, h₂, ← nnrpow_eq_rpow h₃, sqrt_nnrpow (A := A)] diff --git a/Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean b/Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean index d8f00d79a6694..5d627dc71ac06 100644 --- a/Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean +++ b/Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean @@ -330,7 +330,7 @@ theorem Gamma_nat_eq_factorial (n : ℕ) : Gamma (n + 1) = n ! := by @[simp] theorem Gamma_ofNat_eq_factorial (n : ℕ) [(n + 1).AtLeastTwo] : - Gamma (no_index (OfNat.ofNat (n + 1) : ℂ)) = n ! := + Gamma (ofNat(n + 1) : ℂ) = n ! := mod_cast Gamma_nat_eq_factorial (n : ℕ) /-- At `0` the Gamma function is undefined; by convention we assign it the value `0`. -/ @@ -434,7 +434,7 @@ theorem Gamma_nat_eq_factorial (n : ℕ) : Gamma (n + 1) = n ! := by @[simp] theorem Gamma_ofNat_eq_factorial (n : ℕ) [(n + 1).AtLeastTwo] : - Gamma (no_index (OfNat.ofNat (n + 1) : ℝ)) = n ! := + Gamma (ofNat(n + 1) : ℝ) = n ! := mod_cast Gamma_nat_eq_factorial (n : ℕ) /-- At `0` the Gamma function is undefined; by convention we assign it the value `0`. -/ diff --git a/Mathlib/Analysis/SpecialFunctions/Gaussian/FourierTransform.lean b/Mathlib/Analysis/SpecialFunctions/Gaussian/FourierTransform.lean index 89db8290e3687..ee0f874953f99 100644 --- a/Mathlib/Analysis/SpecialFunctions/Gaussian/FourierTransform.lean +++ b/Mathlib/Analysis/SpecialFunctions/Gaussian/FourierTransform.lean @@ -210,9 +210,6 @@ theorem _root_.fourierIntegral_gaussian (hb : 0 < b.re) (t : ℂ) : rw [integral_cexp_quadratic (show (-b).re < 0 by rwa [neg_re, neg_lt_zero]), neg_neg, zero_sub, mul_neg, div_neg, neg_neg, mul_pow, I_sq, neg_one_mul, mul_comm] -@[deprecated (since := "2024-02-21")] -alias _root_.fourier_transform_gaussian := fourierIntegral_gaussian - theorem _root_.fourierIntegral_gaussian_pi' (hb : 0 < b.re) (c : ℂ) : (𝓕 fun x : ℝ => cexp (-π * b * x ^ 2 + 2 * π * c * x)) = fun t : ℝ => 1 / b ^ (1 / 2 : ℂ) * cexp (-π / b * (t + I * c) ^ 2) := by @@ -233,17 +230,11 @@ theorem _root_.fourierIntegral_gaussian_pi' (hb : 0 < b.re) (c : ℂ) : simp only [I_sq] ring -@[deprecated (since := "2024-02-21")] -alias _root_.fourier_transform_gaussian_pi' := _root_.fourierIntegral_gaussian_pi' - theorem _root_.fourierIntegral_gaussian_pi (hb : 0 < b.re) : (𝓕 fun (x : ℝ) ↦ cexp (-π * b * x ^ 2)) = fun t : ℝ ↦ 1 / b ^ (1 / 2 : ℂ) * cexp (-π / b * t ^ 2) := by simpa only [mul_zero, zero_mul, add_zero] using fourierIntegral_gaussian_pi' hb 0 -@[deprecated (since := "2024-02-21")] -alias root_.fourier_transform_gaussian_pi := _root_.fourierIntegral_gaussian_pi - section InnerProductSpace variable {V : Type*} [NormedAddCommGroup V] [InnerProductSpace ℝ V] [FiniteDimensional ℝ V] diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean b/Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean index 43bf1bcc6be45..cbe2f010c0703 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean @@ -176,7 +176,7 @@ alias rpow_nat_cast := rpow_natCast @[simp] lemma rpow_ofNat (x : ℝ≥0) (n : ℕ) [n.AtLeastTwo] : - x ^ (no_index (OfNat.ofNat n) : ℝ) = x ^ (OfNat.ofNat n : ℕ) := + x ^ (ofNat(n) : ℝ) = x ^ (OfNat.ofNat n : ℕ) := rpow_natCast x n theorem rpow_two (x : ℝ≥0) : x ^ (2 : ℝ) = x ^ 2 := rpow_ofNat x 2 @@ -635,7 +635,7 @@ alias rpow_nat_cast := rpow_natCast @[simp] lemma rpow_ofNat (x : ℝ≥0∞) (n : ℕ) [n.AtLeastTwo] : - x ^ (no_index (OfNat.ofNat n) : ℝ) = x ^ (OfNat.ofNat n) := + x ^ (ofNat(n) : ℝ) = x ^ (OfNat.ofNat n) := rpow_natCast x n @[simp, norm_cast] diff --git a/Mathlib/Analysis/SpecificLimits/Basic.lean b/Mathlib/Analysis/SpecificLimits/Basic.lean index c226f8f39d7af..40a14b4c0240c 100644 --- a/Mathlib/Analysis/SpecificLimits/Basic.lean +++ b/Mathlib/Analysis/SpecificLimits/Basic.lean @@ -109,7 +109,7 @@ theorem Filter.EventuallyEq.div_mul_cancel_atTop {α K : Type*} [LinearOrderedSe {f g : α → K} {l : Filter α} (hg : Tendsto g l atTop) : (fun x ↦ f x / g x * g x) =ᶠ[l] fun x ↦ f x := div_mul_cancel <| hg.mono_right <| le_principal_iff.mpr <| - mem_of_superset (Ioi_mem_atTop 0) <| by aesop + mem_of_superset (Ioi_mem_atTop 0) <| by simp /-- If when `x` tends to `∞`, `g` tends to `∞` and `f x / g x` tends to a positive constant, then `f` tends to `∞`. -/ diff --git a/Mathlib/CategoryTheory/Abelian/Exact.lean b/Mathlib/CategoryTheory/Abelian/Exact.lean index badc3ad7a0d10..30b6b7aab74df 100644 --- a/Mathlib/CategoryTheory/Abelian/Exact.lean +++ b/Mathlib/CategoryTheory/Abelian/Exact.lean @@ -105,7 +105,7 @@ def Exact.isLimitImage (h : S.Exact) : rw [exact_iff_kernel_ι_comp_cokernel_π_zero] at h exact KernelFork.IsLimit.ofι _ _ (fun u hu ↦ kernel.lift (cokernel.π S.f) u - (by rw [← kernel.lift_ι S.g u hu, Category.assoc, h, comp_zero])) (by aesop_cat) + (by rw [← kernel.lift_ι S.g u hu, Category.assoc, h, comp_zero])) (by simp) (fun _ _ _ hm => by rw [← cancel_mono (Abelian.image.ι S.f), hm, kernel.lift_ι]) /-- If `(f, g)` is exact, then `image.ι f` is a kernel of `g`. -/ @@ -123,7 +123,7 @@ def Exact.isColimitCoimage (h : S.Exact) : refine CokernelCofork.IsColimit.ofπ _ _ (fun u hu => cokernel.desc (kernel.ι S.g) u (by rw [← cokernel.π_desc S.f u hu, ← Category.assoc, h, zero_comp])) - (by aesop_cat) ?_ + (by simp) ?_ intros _ _ _ _ hm ext rw [hm, cokernel.π_desc] diff --git a/Mathlib/CategoryTheory/Abelian/InjectiveResolution.lean b/Mathlib/CategoryTheory/Abelian/InjectiveResolution.lean index 1473b0ef27413..9110373e9c67a 100644 --- a/Mathlib/CategoryTheory/Abelian/InjectiveResolution.lean +++ b/Mathlib/CategoryTheory/Abelian/InjectiveResolution.lean @@ -240,7 +240,7 @@ lemma InjectiveResolution.iso_hom_naturality {X Y : C} (f : X ⟶ Y) I.iso.hom ≫ (HomotopyCategory.quotient _ _).map φ := by apply HomotopyCategory.eq_of_homotopy apply descHomotopy f - all_goals aesop_cat + all_goals aesop @[reassoc] lemma InjectiveResolution.iso_inv_naturality {X Y : C} (f : X ⟶ Y) diff --git a/Mathlib/CategoryTheory/Action/Limits.lean b/Mathlib/CategoryTheory/Action/Limits.lean index 00fbf901e87fc..a997429802737 100644 --- a/Mathlib/CategoryTheory/Action/Limits.lean +++ b/Mathlib/CategoryTheory/Action/Limits.lean @@ -209,7 +209,7 @@ variable [HasZeroMorphisms V] -- Porting note (https://github.com/leanprover-community/mathlib4/issues/10688): in order to ease automation, the `Zero` instance is introduced separately, -- and the lemma `Action.zero_hom` was moved just below -instance {X Y : Action V G} : Zero (X ⟶ Y) := ⟨0, by aesop_cat⟩ +instance {X Y : Action V G} : Zero (X ⟶ Y) := ⟨0, by simp⟩ @[simp] theorem zero_hom {X Y : Action V G} : (0 : X ⟶ Y).hom = 0 := diff --git a/Mathlib/CategoryTheory/Action/Monoidal.lean b/Mathlib/CategoryTheory/Action/Monoidal.lean index 5428444039ef2..98d9b5620c1a2 100644 --- a/Mathlib/CategoryTheory/Action/Monoidal.lean +++ b/Mathlib/CategoryTheory/Action/Monoidal.lean @@ -89,7 +89,7 @@ variable [BraidedCategory V] instance : BraidedCategory (Action V G) := braidedCategoryOfFaithful (Action.forget V G) (fun X Y => mkIso (β_ _ _) - (fun g => by simp [FunctorCategoryEquivalence.inverse])) (by aesop_cat) + (fun g => by simp [FunctorCategoryEquivalence.inverse])) (by simp) /-- When `V` is braided the forgetful functor `Action V G` to `V` is braided. -/ instance : (Action.forget V G).Braided where diff --git a/Mathlib/CategoryTheory/Adjunction/Basic.lean b/Mathlib/CategoryTheory/Adjunction/Basic.lean index 49a73af0c7d52..6c4183721f9f5 100644 --- a/Mathlib/CategoryTheory/Adjunction/Basic.lean +++ b/Mathlib/CategoryTheory/Adjunction/Basic.lean @@ -293,14 +293,14 @@ theorem eq_homEquiv_apply {A : C} {B : D} (f : F.obj A ⟶ B) (g : A ⟶ G.obj B def corepresentableBy (X : C) : (G ⋙ coyoneda.obj (Opposite.op X)).CorepresentableBy (F.obj X) where homEquiv := adj.homEquiv _ _ - homEquiv_comp := by aesop_cat + homEquiv_comp := by simp /-- If `adj : F ⊣ G`, and `Y : D`, then `G.obj Y` represents `X ↦ (F.obj X ⟶ Y)`-/ @[simps] def representableBy (Y : D) : (F.op ⋙ yoneda.obj Y).RepresentableBy (G.obj Y) where homEquiv := (adj.homEquiv _ _).symm - homEquiv_comp := by aesop_cat + homEquiv_comp := by simp end diff --git a/Mathlib/CategoryTheory/Adjunction/Mates.lean b/Mathlib/CategoryTheory/Adjunction/Mates.lean index cd608363c1820..322bbb825e516 100644 --- a/Mathlib/CategoryTheory/Adjunction/Mates.lean +++ b/Mathlib/CategoryTheory/Adjunction/Mates.lean @@ -101,8 +101,7 @@ def mateEquiv : (G ⋙ L₂ ⟶ L₁ ⋙ H) ≃ (R₁ ⋙ G ⟶ H ⋙ R₂) wher /-- A component of a transposed version of the mates correspondence. -/ theorem mateEquiv_counit (α : G ⋙ L₂ ⟶ L₁ ⋙ H) (d : D) : L₂.map ((mateEquiv adj₁ adj₂ α).app _) ≫ adj₂.counit.app _ = - α.app _ ≫ H.map (adj₁.counit.app d) := by - erw [Functor.map_comp]; simp + α.app _ ≫ H.map (adj₁.counit.app d) := by simp /-- A component of a transposed version of the inverse mates correspondence. -/ theorem mateEquiv_counit_symm (α : R₁ ⋙ G ⟶ H ⋙ R₂) (d : D) : diff --git a/Mathlib/CategoryTheory/Bicategory/FunctorBicategory/Oplax.lean b/Mathlib/CategoryTheory/Bicategory/FunctorBicategory/Oplax.lean index 02482d12bfd09..a0058e9be01b0 100644 --- a/Mathlib/CategoryTheory/Bicategory/FunctorBicategory/Oplax.lean +++ b/Mathlib/CategoryTheory/Bicategory/FunctorBicategory/Oplax.lean @@ -50,19 +50,19 @@ def whiskerRight {η θ : F ⟶ G} (Γ : η ⟶ θ) (ι : G ⟶ H) : η ≫ ι -- Porting note: verified that projections are correct and changed @[simps] to @[simps!] @[simps!] def associator (η : F ⟶ G) (θ : G ⟶ H) (ι : H ⟶ I) : (η ≫ θ) ≫ ι ≅ η ≫ θ ≫ ι := - ModificationIso.ofComponents (fun a => α_ (η.app a) (θ.app a) (ι.app a)) (by aesop_cat) + ModificationIso.ofComponents (fun a => α_ (η.app a) (θ.app a) (ι.app a)) (by simp) /-- Left unitor for the vertical composition of oplax natural transformations. -/ -- Porting note: verified that projections are correct and changed @[simps] to @[simps!] @[simps!] def leftUnitor (η : F ⟶ G) : 𝟙 F ≫ η ≅ η := - ModificationIso.ofComponents (fun a => λ_ (η.app a)) (by aesop_cat) + ModificationIso.ofComponents (fun a => λ_ (η.app a)) (by simp) /-- Right unitor for the vertical composition of oplax natural transformations. -/ -- Porting note: verified that projections are correct and changed @[simps] to @[simps!] @[simps!] def rightUnitor (η : F ⟶ G) : η ≫ 𝟙 G ≅ η := - ModificationIso.ofComponents (fun a => ρ_ (η.app a)) (by aesop_cat) + ModificationIso.ofComponents (fun a => ρ_ (η.app a)) (by simp) end OplaxNatTrans diff --git a/Mathlib/CategoryTheory/Category/Cat.lean b/Mathlib/CategoryTheory/Category/Cat.lean index 7d16c775e8b1d..b76adff0b2f61 100644 --- a/Mathlib/CategoryTheory/Category/Cat.lean +++ b/Mathlib/CategoryTheory/Category/Cat.lean @@ -182,7 +182,7 @@ def typeToCat : Type u ⥤ Cat where simp only [id_eq, eqToHom_refl, Cat.id_map, Category.comp_id, Category.id_comp] apply ULift.ext aesop_cat - · aesop_cat + · simp map_comp f g := by apply Functor.ext; aesop_cat instance : Functor.Faithful typeToCat.{u} where diff --git a/Mathlib/CategoryTheory/Category/Cat/Limit.lean b/Mathlib/CategoryTheory/Category/Cat/Limit.lean index e16a6957e1e3f..ae7ce8e214d0b 100644 --- a/Mathlib/CategoryTheory/Category/Cat/Limit.lean +++ b/Mathlib/CategoryTheory/Category/Cat/Limit.lean @@ -67,10 +67,10 @@ instance (F : J ⥤ Cat.{v, v}) : Category (limit (F ⋙ Cat.objects)) where ← congr_fun (limit.w (homDiagram Y Z) h) g] id_comp _ := by apply Types.limit_ext.{v, v} - aesop_cat + simp comp_id _ := by apply Types.limit_ext.{v, v} - aesop_cat + simp /-- Auxiliary definition: the limit category. -/ @[simps] diff --git a/Mathlib/CategoryTheory/Comma/Basic.lean b/Mathlib/CategoryTheory/Comma/Basic.lean index ccbfebe05bd1b..ee0bc05acbe47 100644 --- a/Mathlib/CategoryTheory/Comma/Basic.lean +++ b/Mathlib/CategoryTheory/Comma/Basic.lean @@ -289,7 +289,7 @@ theorem map_fst : map α β ⋙ fst L' R' = fst L R ⋙ F₁ := where `α : F₁ ⋙ L' ⟶ L ⋙ F`. -/ @[simps!] def mapFst : map α β ⋙ fst L' R' ≅ fst L R ⋙ F₁ := - NatIso.ofComponents (fun _ => Iso.refl _) (by aesop_cat) + NatIso.ofComponents (fun _ => Iso.refl _) (by simp) /-- The equality between `map α β ⋙ snd L' R'` and `snd L R ⋙ F₂`, where `β : R ⋙ F ⟶ F₂ ⋙ R'`. -/ @@ -301,7 +301,7 @@ theorem map_snd : map α β ⋙ snd L' R' = snd L R ⋙ F₂ := where `β : R ⋙ F ⟶ F₂ ⋙ R'`. -/ @[simps!] def mapSnd : map α β ⋙ snd L' R' ≅ snd L R ⋙ F₂ := - NatIso.ofComponents (fun _ => Iso.refl _) (by aesop_cat) + NatIso.ofComponents (fun _ => Iso.refl _) (by simp) end diff --git a/Mathlib/CategoryTheory/Comma/Presheaf/Basic.lean b/Mathlib/CategoryTheory/Comma/Presheaf/Basic.lean index 56df6eace290c..94aa947d47cc9 100644 --- a/Mathlib/CategoryTheory/Comma/Presheaf/Basic.lean +++ b/Mathlib/CategoryTheory/Comma/Presheaf/Basic.lean @@ -190,7 +190,7 @@ lemma yonedaArrow_val {Y : C} {η : yoneda.obj Y ⟶ A} {X : C} {s : yoneda.obj /-- If `η` is also `yoneda`-costructured, then `OverArrows η s` is just morphisms of costructured arrows. -/ def costructuredArrowIso (s t : CostructuredArrow yoneda A) : OverArrows s.hom t.hom ≅ t ⟶ s where - hom p := CostructuredArrow.homMk p.val (by aesop_cat) + hom p := CostructuredArrow.homMk p.val (by simp) inv f := yonedaArrow f.left f.w end OverArrows @@ -436,7 +436,7 @@ lemma app_unitForward {F : Cᵒᵖ ⥤ Type v} (η : F ⟶ A) (X : Cᵒᵖ) /-- Backward direction of the unit. -/ def unitBackward {F : Cᵒᵖ ⥤ Type v} (η : F ⟶ A) (X : C) : F.obj (op X) → YonedaCollection (restrictedYonedaObj η) X := - fun x => YonedaCollection.mk (yonedaEquiv.symm (η.app _ x)) ⟨x, ⟨by aesop_cat⟩⟩ + fun x => YonedaCollection.mk (yonedaEquiv.symm (η.app _ x)) ⟨x, ⟨by simp⟩⟩ lemma unitForward_unitBackward {F : Cᵒᵖ ⥤ Type v} (η : F ⟶ A) (X : C) : unitForward η X ∘ unitBackward η X = id := @@ -513,7 +513,7 @@ lemma counitForward_naturality₂ (s t : (CostructuredArrow yoneda A)ᵒᵖ) (f have : (CostructuredArrow.mkPrecomp t.unop.hom f.unop.left).op = f ≫ eqToHom (by simp [← CostructuredArrow.eq_mk]) := by apply Quiver.Hom.unop_inj - aesop_cat + simp aesop_cat /-- Backward direction of the counit. -/ diff --git a/Mathlib/CategoryTheory/Comma/StructuredArrow/Basic.lean b/Mathlib/CategoryTheory/Comma/StructuredArrow/Basic.lean index 22de055a81308..63d88f4bb3234 100644 --- a/Mathlib/CategoryTheory/Comma/StructuredArrow/Basic.lean +++ b/Mathlib/CategoryTheory/Comma/StructuredArrow/Basic.lean @@ -126,7 +126,7 @@ lemma homMk'_id (f : StructuredArrow S T) : homMk' f (𝟙 f.right) = eqToHom (b ext simp [eqToHom_right] -lemma homMk'_mk_id (f : S ⟶ T.obj Y) : homMk' (mk f) (𝟙 Y) = eqToHom (by aesop_cat) := +lemma homMk'_mk_id (f : S ⟶ T.obj Y) : homMk' (mk f) (𝟙 Y) = eqToHom (by simp) := homMk'_id _ lemma homMk'_comp (f : StructuredArrow S T) (g : f.right ⟶ Y') (g' : Y' ⟶ Y'') : @@ -144,10 +144,10 @@ def mkPostcomp (f : S ⟶ T.obj Y) (g : Y ⟶ Y') : mk f ⟶ mk (f ≫ T.map g) left := 𝟙 _ right := g -lemma mkPostcomp_id (f : S ⟶ T.obj Y) : mkPostcomp f (𝟙 Y) = eqToHom (by aesop_cat) := by aesop_cat +lemma mkPostcomp_id (f : S ⟶ T.obj Y) : mkPostcomp f (𝟙 Y) = eqToHom (by simp) := by simp lemma mkPostcomp_comp (f : S ⟶ T.obj Y) (g : Y ⟶ Y') (g' : Y' ⟶ Y'') : mkPostcomp f (g ≫ g') = mkPostcomp f g ≫ mkPostcomp (f ≫ T.map g) g' ≫ eqToHom (by simp) := by - aesop_cat + simp /-- To construct an isomorphism of structured arrows, we need an isomorphism of the objects underlying the target, @@ -257,7 +257,7 @@ noncomputable def mkIdInitial [T.Full] [T.Faithful] : IsInitial (mk (𝟙 (T.obj desc c := homMk (T.preimage c.pt.hom) uniq c m _ := by apply CommaMorphism.ext - · aesop_cat + · simp · apply T.map_injective simpa only [homMk_right, T.map_preimage, ← w m] using (Category.id_comp _).symm @@ -293,7 +293,7 @@ instance (S : C) (F : B ⥤ C) (G : C ⥤ D) : (post S F G).Faithful where map_injective {_ _} _ _ h := by simpa [ext_iff] using h instance (S : C) (F : B ⥤ C) (G : C ⥤ D) [G.Faithful] : (post S F G).Full where - map_surjective f := ⟨homMk f.right (G.map_injective (by simpa using f.w.symm)), by aesop_cat⟩ + map_surjective f := ⟨homMk f.right (G.map_injective (by simpa using f.w.symm)), by simp⟩ instance (S : C) (F : B ⥤ C) (G : C ⥤ D) [G.Full] : (post S F G).EssSurj where mem_essImage h := ⟨mk (G.preimage h.hom), ⟨isoMk (Iso.refl _) (by simp)⟩⟩ @@ -468,7 +468,7 @@ lemma homMk'_id (f : CostructuredArrow S T) : homMk' f (𝟙 f.left) = eqToHom ( ext simp [eqToHom_left] -lemma homMk'_mk_id (f : S.obj Y ⟶ T) : homMk' (mk f) (𝟙 Y) = eqToHom (by aesop_cat) := +lemma homMk'_mk_id (f : S.obj Y ⟶ T) : homMk' (mk f) (𝟙 Y) = eqToHom (by simp) := homMk'_id _ lemma homMk'_comp (f : CostructuredArrow S T) (g : Y' ⟶ f.left) (g' : Y'' ⟶ Y') : @@ -486,10 +486,10 @@ def mkPrecomp (f : S.obj Y ⟶ T) (g : Y' ⟶ Y) : mk (S.map g ≫ f) ⟶ mk f w left := g right := 𝟙 _ -lemma mkPrecomp_id (f : S.obj Y ⟶ T) : mkPrecomp f (𝟙 Y) = eqToHom (by aesop_cat) := by aesop_cat +lemma mkPrecomp_id (f : S.obj Y ⟶ T) : mkPrecomp f (𝟙 Y) = eqToHom (by simp) := by simp lemma mkPrecomp_comp (f : S.obj Y ⟶ T) (g : Y' ⟶ Y) (g' : Y'' ⟶ Y') : mkPrecomp f (g' ≫ g) = eqToHom (by simp) ≫ mkPrecomp (S.map g ≫ f) g' ≫ mkPrecomp f g := by - aesop_cat + simp /-- To construct an isomorphism of costructured arrows, we need an isomorphism of the objects underlying the source, @@ -634,7 +634,7 @@ instance (F : B ⥤ C) (G : C ⥤ D) (S : C) : (post F G S).Faithful where map_injective {_ _} _ _ h := by simpa [ext_iff] using h instance (F : B ⥤ C) (G : C ⥤ D) (S : C) [G.Faithful] : (post F G S).Full where - map_surjective f := ⟨homMk f.left (G.map_injective (by simpa using f.w)), by aesop_cat⟩ + map_surjective f := ⟨homMk f.left (G.map_injective (by simpa using f.w)), by simp⟩ instance (F : B ⥤ C) (G : C ⥤ D) (S : C) [G.Full] : (post F G S).EssSurj where mem_essImage h := ⟨mk (G.preimage h.hom), ⟨isoMk (Iso.refl _) (by simp)⟩⟩ @@ -783,12 +783,7 @@ category of structured arrows `d ⟶ F.obj c` to the category of costructured ar def toCostructuredArrow (F : C ⥤ D) (d : D) : (StructuredArrow d F)ᵒᵖ ⥤ CostructuredArrow F.op (op d) where obj X := @CostructuredArrow.mk _ _ _ _ _ (op X.unop.right) F.op X.unop.hom.op - map f := - CostructuredArrow.homMk f.unop.right.op - (by - dsimp - rw [← op_comp, ← f.unop.w, Functor.const_obj_map] - erw [Category.id_comp]) + map f := CostructuredArrow.homMk f.unop.right.op (by simp [← op_comp]) /-- For a functor `F : C ⥤ D` and an object `d : D`, we obtain a contravariant functor from the category of structured arrows `op d ⟶ F.op.obj c` to the category of costructured arrows @@ -803,8 +798,8 @@ def toCostructuredArrow' (F : C ⥤ D) (d : D) : (by dsimp rw [← Quiver.Hom.unop_op (F.map (Quiver.Hom.unop f.unop.right)), ← unop_comp, ← F.op_map, ← - f.unop.w, Functor.const_obj_map] - erw [Category.id_comp]) + f.unop.w] + simp) end StructuredArrow @@ -818,12 +813,7 @@ category of costructured arrows `F.obj c ⟶ d` to the category of structured ar def toStructuredArrow (F : C ⥤ D) (d : D) : (CostructuredArrow F d)ᵒᵖ ⥤ StructuredArrow (op d) F.op where obj X := @StructuredArrow.mk _ _ _ _ _ (op X.unop.left) F.op X.unop.hom.op - map f := - StructuredArrow.homMk f.unop.left.op - (by - dsimp - rw [← op_comp, f.unop.w, Functor.const_obj_map] - erw [Category.comp_id]) + map f := StructuredArrow.homMk f.unop.left.op (by simp [← op_comp]) /-- For a functor `F : C ⥤ D` and an object `d : D`, we obtain a contravariant functor from the category of costructured arrows `F.op.obj c ⟶ op d` to the category of structured arrows @@ -839,7 +829,7 @@ def toStructuredArrow' (F : C ⥤ D) (d : D) : dsimp rw [← Quiver.Hom.unop_op (F.map f.unop.left.unop), ← unop_comp, ← F.op_map, f.unop.w, Functor.const_obj_map] - erw [Category.comp_id]) + simp) end CostructuredArrow diff --git a/Mathlib/CategoryTheory/ComposableArrows.lean b/Mathlib/CategoryTheory/ComposableArrows.lean index 45cd88fc81f69..9d9c9bb7cdc6d 100644 --- a/Mathlib/CategoryTheory/ComposableArrows.lean +++ b/Mathlib/CategoryTheory/ComposableArrows.lean @@ -335,10 +335,7 @@ lemma map_one_succ (j : ℕ) (hj : j + 1 < n + 1 + 1) : map F f 1 ⟨j + 1, hj⟩ (by simp [Fin.le_def]) = F.map' 0 j := rfl lemma map_id (i : Fin (n + 1 + 1)) : map F f i i (by simp) = 𝟙 _ := by - obtain ⟨i, hi⟩ := i - cases i - · rfl - · apply F.map_id + obtain ⟨_|_, hi⟩ := i <;> simp lemma map_comp {i j k : Fin (n + 1 + 1)} (hij : i ≤ j) (hjk : j ≤ k) : map F f i k (hij.trans hjk) = map F f i j hij ≫ map F f j k hjk := by @@ -542,14 +539,8 @@ lemma ext_succ {F G : ComposableArrows C (n + 1)} (h₀ : F.obj' 0 = G.obj' 0) exact Functor.ext_of_iso (isoMkSucc (eqToIso h₀) (eqToIso h) (by rw [w] dsimp [app'] - rw [eqToHom_app, assoc, assoc, eqToHom_trans, eqToHom_refl, comp_id])) this (by - rintro ⟨i, hi⟩ - dsimp - cases' i with i - · erw [homMkSucc_app_zero] - · rw [homMkSucc_app_succ] - dsimp [app'] - rw [eqToHom_app]) + rw [eqToHom_app, assoc, assoc, eqToHom_trans, eqToHom_refl, comp_id])) this + (by rintro ⟨_|_, hi⟩ <;> simp) lemma precomp_surjective (F : ComposableArrows C (n + 1)) : ∃ (F₀ : ComposableArrows C n) (X₀ : C) (f₀ : X₀ ⟶ F₀.left), F = F₀.precomp f₀ := diff --git a/Mathlib/CategoryTheory/DiscreteCategory.lean b/Mathlib/CategoryTheory/DiscreteCategory.lean index 1d28f0410aa9e..4505a8e1b9584 100644 --- a/Mathlib/CategoryTheory/DiscreteCategory.lean +++ b/Mathlib/CategoryTheory/DiscreteCategory.lean @@ -246,9 +246,9 @@ def equivalence {I : Type u₁} {J : Type u₂} (e : I ≃ J) : Discrete I ≌ D functor := Discrete.functor (Discrete.mk ∘ (e : I → J)) inverse := Discrete.functor (Discrete.mk ∘ (e.symm : J → I)) unitIso := - Discrete.natIso fun i => eqToIso (by aesop_cat) + Discrete.natIso fun i => eqToIso (by simp) counitIso := - Discrete.natIso fun j => eqToIso (by aesop_cat) + Discrete.natIso fun j => eqToIso (by simp) /-- We can convert an equivalence of `discrete` categories to a type-level `Equiv`. -/ @[simps] diff --git a/Mathlib/CategoryTheory/EffectiveEpi/Coproduct.lean b/Mathlib/CategoryTheory/EffectiveEpi/Coproduct.lean index 80a53351c358d..87c5560d945e1 100644 --- a/Mathlib/CategoryTheory/EffectiveEpi/Coproduct.lean +++ b/Mathlib/CategoryTheory/EffectiveEpi/Coproduct.lean @@ -36,20 +36,12 @@ def effectiveEpiStructIsColimitDescOfEffectiveEpiFamily {B : C} {α : Type*} (X uniq e _ m hm := EffectiveEpiFamily.uniq X π (fun a ↦ c.ι.app ⟨a⟩ ≫ e) (fun _ _ _ _ hg ↦ (by simp [← hm, reassoc_of% hg])) m (fun _ ↦ (by simp [← hm])) -/-- -Given an `EffectiveEpiFamily X π` such that the coproduct of `X` exists, `Sigma.desc π` is an -`EffectiveEpi`. --/ -noncomputable -def effectiveEpiStructDescOfEffectiveEpiFamily {B : C} {α : Type*} (X : α → C) - (π : (a : α) → (X a ⟶ B)) [HasCoproduct X] [EffectiveEpiFamily X π] : - EffectiveEpiStruct (Sigma.desc π) := by - simpa [coproductIsCoproduct] using - effectiveEpiStructIsColimitDescOfEffectiveEpiFamily X _ (coproductIsCoproduct _) π - instance {B : C} {α : Type*} (X : α → C) (π : (a : α) → (X a ⟶ B)) [HasCoproduct X] - [EffectiveEpiFamily X π] : EffectiveEpi (Sigma.desc π) := - ⟨⟨effectiveEpiStructDescOfEffectiveEpiFamily X π⟩⟩ + [EffectiveEpiFamily X π] : EffectiveEpi (Sigma.desc π) := by + let e := effectiveEpiStructIsColimitDescOfEffectiveEpiFamily X _ (coproductIsCoproduct _) π + simp only [Cofan.mk_pt, coproductIsCoproduct, colimit.cocone_x, IsColimit.ofIsoColimit_desc, + Cocones.ext_inv_hom, Iso.refl_inv, colimit.isColimit_desc, Category.id_comp] at e + exact ⟨⟨e⟩⟩ example {B : C} {α : Type*} (X : α → C) (π : (a : α) → (X a ⟶ B)) [EffectiveEpiFamily X π] [HasCoproduct X] : Epi (Sigma.desc π) := inferInstance diff --git a/Mathlib/CategoryTheory/Elements.lean b/Mathlib/CategoryTheory/Elements.lean index 7d85416b30755..01aa8d0ac97b7 100644 --- a/Mathlib/CategoryTheory/Elements.lean +++ b/Mathlib/CategoryTheory/Elements.lean @@ -59,7 +59,7 @@ lemma Functor.Elements.ext {F : C ⥤ Type w} (x y : F.Elements) (h₁ : x.fst = -/ instance categoryOfElements (F : C ⥤ Type w) : Category.{v} F.Elements where Hom p q := { f : p.1 ⟶ q.1 // (F.map f) p.2 = q.2 } - id p := ⟨𝟙 p.1, by aesop_cat⟩ + id p := ⟨𝟙 p.1, by simp⟩ comp {X Y Z} f g := ⟨f.val ≫ g.val, by simp [f.2, g.2]⟩ /-- Natural transformations are mapped to functors between category of elements -/ diff --git a/Mathlib/CategoryTheory/Enriched/Basic.lean b/Mathlib/CategoryTheory/Enriched/Basic.lean index bd179b0fdb051..0654c73c33d3c 100644 --- a/Mathlib/CategoryTheory/Enriched/Basic.lean +++ b/Mathlib/CategoryTheory/Enriched/Basic.lean @@ -143,7 +143,7 @@ def categoryOfEnrichedCategoryType (C : Type u₁) [𝒞 : EnrichedCategory (Typ comp f g := eComp (Type v) _ _ _ ⟨f, g⟩ id_comp f := congr_fun (e_id_comp (Type v) _ _) f comp_id f := congr_fun (e_comp_id (Type v) _ _) f - assoc f g h := (congr_fun (e_assoc (Type v) _ _ _ _) ⟨f, g, h⟩ : _) + assoc f g h := (congr_fun (e_assoc (Type v) _ _ _ _) ⟨f, g, h⟩ :) /-- Construct a `Type v`-enriched category from an honest category. -/ diff --git a/Mathlib/CategoryTheory/Equivalence.lean b/Mathlib/CategoryTheory/Equivalence.lean index 949a01797a9d6..85a5fc5503a8c 100644 --- a/Mathlib/CategoryTheory/Equivalence.lean +++ b/Mathlib/CategoryTheory/Equivalence.lean @@ -182,7 +182,7 @@ theorem unit_inverse_comp (e : C ≌ D) (Y : D) : slice_lhs 2 3 => erw [← map_comp e.inverse, ← e.counitIso.inv.naturality, (e.counitIso.app _).hom_inv_id, map_id] - erw [id_comp, (e.unitIso.app _).hom_inv_id]; rfl + simp @[reassoc (attr := simp)] theorem inverse_counitInv_comp (e : C ≌ D) (Y : D) : @@ -292,13 +292,13 @@ def funInvIdAssoc (e : C ≌ D) (F : C ⥤ E) : e.functor ⋙ e.inverse ⋙ F theorem funInvIdAssoc_hom_app (e : C ≌ D) (F : C ⥤ E) (X : C) : (funInvIdAssoc e F).hom.app X = F.map (e.unitInv.app X) := by dsimp [funInvIdAssoc] - aesop_cat + simp @[simp] theorem funInvIdAssoc_inv_app (e : C ≌ D) (F : C ⥤ E) (X : C) : (funInvIdAssoc e F).inv.app X = F.map (e.unit.app X) := by dsimp [funInvIdAssoc] - aesop_cat + simp /-- Composing a functor with both functors of an equivalence yields a naturally isomorphic functor. -/ @@ -309,13 +309,13 @@ def invFunIdAssoc (e : C ≌ D) (F : D ⥤ E) : e.inverse ⋙ e.functor ⋙ F theorem invFunIdAssoc_hom_app (e : C ≌ D) (F : D ⥤ E) (X : D) : (invFunIdAssoc e F).hom.app X = F.map (e.counit.app X) := by dsimp [invFunIdAssoc] - aesop_cat + simp @[simp] theorem invFunIdAssoc_inv_app (e : C ≌ D) (F : D ⥤ E) (X : D) : (invFunIdAssoc e F).inv.app X = F.map (e.counitInv.app X) := by dsimp [invFunIdAssoc] - aesop_cat + simp /-- If `C` is equivalent to `D`, then `C ⥤ E` is equivalent to `D ⥤ E`. -/ @[simps! functor inverse unitIso counitIso] @@ -524,7 +524,7 @@ i.e. faithful, full, and essentially surjective. -/ noncomputable def inv (F : C ⥤ D) [F.IsEquivalence] : D ⥤ C where obj X := F.objPreimage X map {X Y} f := F.preimage ((F.objObjPreimageIso X).hom ≫ f ≫ (F.objObjPreimageIso Y).inv) - map_id X := by apply F.map_injective; aesop_cat + map_id X := by apply F.map_injective; simp map_comp {X Y Z} f g := by apply F.map_injective; simp /-- Interpret a functor that is an equivalence as an equivalence. diff --git a/Mathlib/CategoryTheory/Extensive.lean b/Mathlib/CategoryTheory/Extensive.lean index 18e02517b1d56..f7766dcde513a 100644 --- a/Mathlib/CategoryTheory/Extensive.lean +++ b/Mathlib/CategoryTheory/Extensive.lean @@ -174,10 +174,10 @@ theorem FinitaryExtensive.mono_inl_of_isColimit [FinitaryExtensive C] {c : Binar FinitaryExtensive.mono_inr_of_isColimit (BinaryCofan.isColimitFlip hc) instance [FinitaryExtensive C] (X Y : C) : Mono (coprod.inl : X ⟶ X ⨿ Y) := - (FinitaryExtensive.mono_inl_of_isColimit (coprodIsCoprod X Y) : _) + (FinitaryExtensive.mono_inl_of_isColimit (coprodIsCoprod X Y) :) instance [FinitaryExtensive C] (X Y : C) : Mono (coprod.inr : Y ⟶ X ⨿ Y) := - (FinitaryExtensive.mono_inr_of_isColimit (coprodIsCoprod X Y) : _) + (FinitaryExtensive.mono_inr_of_isColimit (coprodIsCoprod X Y) :) theorem FinitaryExtensive.isPullback_initial_to_binaryCofan [FinitaryExtensive C] {c : BinaryCofan X Y} (hc : IsColimit c) : @@ -222,7 +222,7 @@ instance types.finitaryExtensive : FinitaryExtensive (Type u) := by · simp only [Types.binaryCoproductCocone_pt, Functor.const_obj_obj, Sum.inl.injEq, existsUnique_eq'] · apply_fun f at h - cases ((congr_fun s.condition x).symm.trans h).trans (congr_fun hαY val : _).symm + cases ((congr_fun s.condition x).symm.trans h).trans (congr_fun hαY val :).symm delta ExistsUnique at this choose l hl hl' using this exact ⟨l, (funext hl).symm, Types.isTerminalPunit.hom_ext _ _, @@ -233,7 +233,7 @@ instance types.finitaryExtensive : FinitaryExtensive (Type u) := by intro x cases' h : s.fst x with val val · apply_fun f at h - cases ((congr_fun s.condition x).symm.trans h).trans (congr_fun hαX val : _).symm + cases ((congr_fun s.condition x).symm.trans h).trans (congr_fun hαX val :).symm · simp only [Types.binaryCoproductCocone_pt, Functor.const_obj_obj, Sum.inr.injEq, existsUnique_eq'] delta ExistsUnique at this @@ -314,7 +314,7 @@ instance finitaryExtensive_TopCat : FinitaryExtensive TopCat.{u} := by · exact ⟨val, rfl, fun y h => Sum.inl_injective h.symm⟩ · apply_fun f at h cases ((ConcreteCategory.congr_hom s.condition x).symm.trans h).trans - (ConcreteCategory.congr_hom hαY val : _).symm + (ConcreteCategory.congr_hom hαY val :).symm delta ExistsUnique at this choose l hl hl' using this refine ⟨⟨l, ?_⟩, ContinuousMap.ext fun a => (hl a).symm, TopCat.isTerminalPUnit.hom_ext _ _, @@ -330,7 +330,7 @@ instance finitaryExtensive_TopCat : FinitaryExtensive TopCat.{u} := by cases' h : s.fst x with val val · apply_fun f at h cases ((ConcreteCategory.congr_hom s.condition x).symm.trans h).trans - (ConcreteCategory.congr_hom hαX val : _).symm + (ConcreteCategory.congr_hom hαX val :).symm · exact ⟨val, rfl, fun y h => Sum.inr_injective h.symm⟩ delta ExistsUnique at this choose l hl hl' using this diff --git a/Mathlib/CategoryTheory/Filtered/Basic.lean b/Mathlib/CategoryTheory/Filtered/Basic.lean index 584411d3326a4..e356e98a41ca6 100644 --- a/Mathlib/CategoryTheory/Filtered/Basic.lean +++ b/Mathlib/CategoryTheory/Filtered/Basic.lean @@ -829,9 +829,6 @@ theorem of_isInitial {X : C} (h : IsInitial X) : IsCofiltered C := instance (priority := 100) of_hasInitial [HasInitial C] : IsCofiltered C := of_isInitial _ initialIsInitial -@[deprecated (since := "2024-03-11")] -alias _root_.CategoryTheory.cofiltered_of_hasFiniteLimits := of_hasFiniteLimits - /-- For every universe `w`, `C` is filtered if and only if every finite diagram in `C` with shape in `w` admits a cocone. -/ theorem iff_cone_nonempty : IsCofiltered C ↔ diff --git a/Mathlib/CategoryTheory/Functor/Const.lean b/Mathlib/CategoryTheory/Functor/Const.lean index ba2d6a7b37e98..63504ca0530c7 100644 --- a/Mathlib/CategoryTheory/Functor/Const.lean +++ b/Mathlib/CategoryTheory/Functor/Const.lean @@ -96,7 +96,7 @@ instance [Nonempty J] : Faithful (const J : C ⥤ J ⥤ C) where def compConstIso (F : C ⥤ D) : F ⋙ Functor.const J ≅ Functor.const J ⋙ (whiskeringRight J C D).obj F := NatIso.ofComponents - (fun X => NatIso.ofComponents (fun _ => Iso.refl _) (by aesop_cat)) + (fun X => NatIso.ofComponents (fun _ => Iso.refl _) (by simp)) (by aesop_cat) /-- The canonical isomorphism diff --git a/Mathlib/CategoryTheory/Functor/Derived/RightDerived.lean b/Mathlib/CategoryTheory/Functor/Derived/RightDerived.lean index 8acc96f8e2d2d..c815ff3015b16 100644 --- a/Mathlib/CategoryTheory/Functor/Derived/RightDerived.lean +++ b/Mathlib/CategoryTheory/Functor/Derived/RightDerived.lean @@ -117,7 +117,7 @@ lemma rightDerivedNatTrans_app (τ : F ⟶ F') (X : C) : @[simp] lemma rightDerivedNatTrans_id : rightDerivedNatTrans RF RF α α W (𝟙 F) = 𝟙 RF := - rightDerived_ext RF α W _ _ _ (by aesop_cat) + rightDerived_ext RF α W _ _ _ (by simp) variable [RF'.IsRightDerivedFunctor α' W] @@ -125,7 +125,7 @@ variable [RF'.IsRightDerivedFunctor α' W] lemma rightDerivedNatTrans_comp (τ : F ⟶ F') (τ' : F' ⟶ F'') : rightDerivedNatTrans RF RF' α α' W τ ≫ rightDerivedNatTrans RF' RF'' α' α'' W τ' = rightDerivedNatTrans RF RF'' α α'' W (τ ≫ τ') := - rightDerived_ext RF α W _ _ _ (by aesop_cat) + rightDerived_ext RF α W _ _ _ (by simp) /-- The natural isomorphism `RF ≅ RF'` on right derived functors that is induced by a natural isomorphism `F ≅ F'`. -/ diff --git a/Mathlib/CategoryTheory/Functor/Flat.lean b/Mathlib/CategoryTheory/Functor/Flat.lean index 0cb417530c71d..981c2d3094209 100644 --- a/Mathlib/CategoryTheory/Functor/Flat.lean +++ b/Mathlib/CategoryTheory/Functor/Flat.lean @@ -195,9 +195,9 @@ theorem uniq {K : J ⥤ C} {c : Cone K} (hc : IsLimit c) (s : Cone (K ⋙ F)) let α₂ : (F.mapCone c).toStructuredArrow ⋙ map f₂ ⟶ s.toStructuredArrow := { app := fun X => eqToHom (by simp [← h₂]) } let c₁ : Cone (s.toStructuredArrow ⋙ pre s.pt K F) := - (Cones.postcompose (whiskerRight α₁ (pre s.pt K F) : _)).obj (c.toStructuredArrowCone F f₁) + (Cones.postcompose (whiskerRight α₁ (pre s.pt K F) :)).obj (c.toStructuredArrowCone F f₁) let c₂ : Cone (s.toStructuredArrow ⋙ pre s.pt K F) := - (Cones.postcompose (whiskerRight α₂ (pre s.pt K F) : _)).obj (c.toStructuredArrowCone F f₂) + (Cones.postcompose (whiskerRight α₂ (pre s.pt K F) :)).obj (c.toStructuredArrowCone F f₂) -- The two cones can then be combined and we may obtain a cone over the two cones since -- `StructuredArrow s.pt F` is cofiltered. let c₀ := IsCofiltered.cone (biconeMk _ c₁ c₂) diff --git a/Mathlib/CategoryTheory/Functor/FullyFaithful.lean b/Mathlib/CategoryTheory/Functor/FullyFaithful.lean index 82f3de969be51..c550e1890088a 100644 --- a/Mathlib/CategoryTheory/Functor/FullyFaithful.lean +++ b/Mathlib/CategoryTheory/Functor/FullyFaithful.lean @@ -64,7 +64,7 @@ lemma map_injective_iff (F : C ⥤ D) [Faithful F] {X Y : C} (f g : X ⟶ Y) : theorem mapIso_injective (F : C ⥤ D) [Faithful F] : Function.Injective <| (F.mapIso : (X ≅ Y) → (F.obj X ≅ F.obj Y)) := fun _ _ h => - Iso.ext (map_injective F (congr_arg Iso.hom h : _)) + Iso.ext (map_injective F (congr_arg Iso.hom h :)) theorem map_surjective (F : C ⥤ D) [Full F] : Function.Surjective (F.map : (X ⟶ Y) → (F.obj X ⟶ F.obj Y)) := diff --git a/Mathlib/CategoryTheory/Functor/FunctorHom.lean b/Mathlib/CategoryTheory/Functor/FunctorHom.lean index 3f4312c687eaa..be91d1395df1e 100644 --- a/Mathlib/CategoryTheory/Functor/FunctorHom.lean +++ b/Mathlib/CategoryTheory/Functor/FunctorHom.lean @@ -131,7 +131,7 @@ def functorHomEquiv (A : C ⥤ Type max u v v') : (A ⟶ F.functorHom G) ≃ Hom ext X a Y f exact (HomObj.congr_app (congr_fun (φ.naturality f) a) Y (𝟙 _)).trans (congr_arg ((φ.app X a).app Y) (by simp)) - right_inv x := by aesop + right_inv x := by simp variable {F G} in /-- Morphisms `(𝟙_ (C ⥤ Type max v' v u) ⟶ F.functorHom G)` are in bijection with diff --git a/Mathlib/CategoryTheory/Functor/KanExtension/Basic.lean b/Mathlib/CategoryTheory/Functor/KanExtension/Basic.lean index 4ea68229bb26b..b969b1f888ff4 100644 --- a/Mathlib/CategoryTheory/Functor/KanExtension/Basic.lean +++ b/Mathlib/CategoryTheory/Functor/KanExtension/Basic.lean @@ -102,7 +102,7 @@ noncomputable def homEquivOfIsRightKanExtension (G : D ⥤ H) : (G ⟶ F') ≃ (L ⋙ G ⟶ F) where toFun β := whiskerLeft _ β ≫ α invFun β := liftOfIsRightKanExtension _ α _ β - left_inv β := Functor.hom_ext_of_isRightKanExtension _ α _ _ (by aesop_cat) + left_inv β := Functor.hom_ext_of_isRightKanExtension _ α _ _ (by simp) right_inv := by aesop_cat lemma isRightKanExtension_of_iso {F' F'' : D ⥤ H} (e : F' ≅ F'') {L : C ⥤ D} {F : C ⥤ H} @@ -193,7 +193,7 @@ noncomputable def homEquivOfIsLeftKanExtension (G : D ⥤ H) : (F' ⟶ G) ≃ (F ⟶ L ⋙ G) where toFun β := α ≫ whiskerLeft _ β invFun β := descOfIsLeftKanExtension _ α _ β - left_inv β := Functor.hom_ext_of_isLeftKanExtension _ α _ _ (by aesop_cat) + left_inv β := Functor.hom_ext_of_isLeftKanExtension _ α _ _ (by simp) right_inv := by aesop_cat lemma isLeftKanExtension_of_iso {F' : D ⥤ H} {F'' : D ⥤ H} (e : F' ≅ F'') diff --git a/Mathlib/CategoryTheory/Functor/KanExtension/Pointwise.lean b/Mathlib/CategoryTheory/Functor/KanExtension/Pointwise.lean index a1a044d1a3846..12c6638bcdfe2 100644 --- a/Mathlib/CategoryTheory/Functor/KanExtension/Pointwise.lean +++ b/Mathlib/CategoryTheory/Functor/KanExtension/Pointwise.lean @@ -140,8 +140,8 @@ def isPointwiseLeftKanExtensionEquivOfIso (e : E ≅ E') : E.IsPointwiseLeftKanExtension ≃ E'.IsPointwiseLeftKanExtension where toFun h := fun Y => (isPointwiseLeftKanExtensionAtEquivOfIso e Y) (h Y) invFun h := fun Y => (isPointwiseLeftKanExtensionAtEquivOfIso e Y).symm (h Y) - left_inv h := by aesop - right_inv h := by aesop + left_inv h := by simp + right_inv h := by simp variable (h : E.IsPointwiseLeftKanExtension) include h @@ -277,8 +277,8 @@ def isPointwiseRightKanExtensionEquivOfIso (e : E ≅ E') : E.IsPointwiseRightKanExtension ≃ E'.IsPointwiseRightKanExtension where toFun h := fun Y => (isPointwiseRightKanExtensionAtEquivOfIso e Y) (h Y) invFun h := fun Y => (isPointwiseRightKanExtensionAtEquivOfIso e Y).symm (h Y) - left_inv h := by aesop - right_inv h := by aesop + left_inv h := by simp + right_inv h := by simp variable (h : E.IsPointwiseRightKanExtension) include h diff --git a/Mathlib/CategoryTheory/Galois/Decomposition.lean b/Mathlib/CategoryTheory/Galois/Decomposition.lean index 0266fabf8bc16..37f5ba335107f 100644 --- a/Mathlib/CategoryTheory/Galois/Decomposition.lean +++ b/Mathlib/CategoryTheory/Galois/Decomposition.lean @@ -66,7 +66,7 @@ private lemma has_decomp_connected_components_aux_initial (X : C) (h : IsInitial ∃ (ι : Type) (f : ι → C) (g : (i : ι) → (f i) ⟶ X) (_ : IsColimit (Cofan.mk X g)), (∀ i, IsConnected (f i)) ∧ Finite ι := by refine ⟨Empty, fun _ ↦ X, fun _ ↦ 𝟙 X, ?_⟩ - use mkCofanColimit _ (fun s ↦ IsInitial.to h s.pt) (fun s ↦ by aesop) + use mkCofanColimit _ (fun s ↦ IsInitial.to h s.pt) (fun s ↦ by simp) (fun s m _ ↦ IsInitial.hom_ext h m _) exact ⟨by simp only [IsEmpty.forall_iff], inferInstance⟩ diff --git a/Mathlib/CategoryTheory/GlueData.lean b/Mathlib/CategoryTheory/GlueData.lean index 2c0b603d6f43f..fcea852442c19 100644 --- a/Mathlib/CategoryTheory/GlueData.lean +++ b/Mathlib/CategoryTheory/GlueData.lean @@ -443,9 +443,9 @@ def GlueData'.t'' (D : GlueData' C) (i j k : D.J) : else haveI := Ne.symm hij pullback.map _ _ _ _ (eqToHom (by aesop)) (eqToHom (by rw [dif_neg hik])) - (eqToHom (by aesop)) (by delta f'; aesop) (by delta f'; aesop) ≫ + (eqToHom (by simp)) (by delta f'; aesop) (by delta f'; aesop) ≫ D.t' i j k hij hik hjk ≫ - pullback.map _ _ _ _ (eqToHom (by aesop)) (eqToHom (by aesop)) (eqToHom (by aesop)) + pullback.map _ _ _ _ (eqToHom (by aesop)) (eqToHom (by aesop)) (eqToHom (by simp)) (by delta f'; aesop) (by delta f'; aesop) open scoped Classical in diff --git a/Mathlib/CategoryTheory/GradedObject/Single.lean b/Mathlib/CategoryTheory/GradedObject/Single.lean index 180a200b1ed92..e5add88c366ef 100644 --- a/Mathlib/CategoryTheory/GradedObject/Single.lean +++ b/Mathlib/CategoryTheory/GradedObject/Single.lean @@ -77,7 +77,7 @@ variable (C) in evaluation functor `eval j` identifies to the identity functor. -/ @[simps!] noncomputable def singleCompEval (j : J) : single j ⋙ eval j ≅ 𝟭 C := - NatIso.ofComponents (singleObjApplyIso j) (by aesop_cat) + NatIso.ofComponents (singleObjApplyIso j) (by simp) end GradedObject diff --git a/Mathlib/CategoryTheory/Grothendieck.lean b/Mathlib/CategoryTheory/Grothendieck.lean index 2d28b0e705e8b..c987c94e3af65 100644 --- a/Mathlib/CategoryTheory/Grothendieck.lean +++ b/Mathlib/CategoryTheory/Grothendieck.lean @@ -87,7 +87,7 @@ theorem ext {X Y : Grothendieck F} (f g : Hom X Y) (w_base : f.base = g.base) -/ def id (X : Grothendieck F) : Hom X X where base := 𝟙 X.base - fiber := eqToHom (by erw [CategoryTheory.Functor.map_id, Functor.id_obj X.fiber]) + fiber := eqToHom (by simp) instance (X : Grothendieck F) : Inhabited (Hom X X) := ⟨id X⟩ @@ -97,7 +97,7 @@ instance (X : Grothendieck F) : Inhabited (Hom X X) := def comp {X Y Z : Grothendieck F} (f : Hom X Y) (g : Hom Y Z) : Hom X Z where base := f.base ≫ g.base fiber := - eqToHom (by erw [Functor.map_comp, Functor.comp_obj]) ≫ (F.map g.base).map f.fiber ≫ g.fiber + eqToHom (by simp) ≫ (F.map g.base).map f.fiber ≫ g.fiber attribute [local simp] eqToHom_map @@ -126,7 +126,7 @@ theorem id_base (X : Grothendieck F) : @[simp] theorem id_fiber (X : Grothendieck F) : - Hom.fiber (𝟙 X) = eqToHom (by erw [CategoryTheory.Functor.map_id, Functor.id_obj X.fiber]) := + Hom.fiber (𝟙 X) = eqToHom (by simp) := rfl @[simp] @@ -137,11 +137,9 @@ theorem comp_base {X Y Z : Grothendieck F} (f : X ⟶ Y) (g : Y ⟶ Z) : @[simp] theorem comp_fiber {X Y Z : Grothendieck F} (f : X ⟶ Y) (g : Y ⟶ Z) : Hom.fiber (f ≫ g) = - eqToHom (by erw [Functor.map_comp, Functor.comp_obj]) ≫ - (F.map g.base).map f.fiber ≫ g.fiber := + eqToHom (by simp) ≫ (F.map g.base).map f.fiber ≫ g.fiber := rfl - theorem congr {X Y : Grothendieck F} {f g : X ⟶ Y} (h : f = g) : f.fiber = eqToHom (by subst h; rfl) ≫ g.fiber := by subst h diff --git a/Mathlib/CategoryTheory/Groupoid.lean b/Mathlib/CategoryTheory/Groupoid.lean index a880d7c21b5a3..5b43b7b7a2a4b 100644 --- a/Mathlib/CategoryTheory/Groupoid.lean +++ b/Mathlib/CategoryTheory/Groupoid.lean @@ -95,7 +95,7 @@ variable (X Y) /-- In a groupoid, isomorphisms are equivalent to morphisms. -/ def Groupoid.isoEquivHom : (X ≅ Y) ≃ (X ⟶ Y) where toFun := Iso.hom - invFun f := ⟨f, Groupoid.inv f, (by aesop_cat), (by aesop_cat)⟩ + invFun f := ⟨f, Groupoid.inv f, (by simp), (by simp)⟩ left_inv _ := Iso.ext rfl right_inv _ := rfl diff --git a/Mathlib/CategoryTheory/Idempotents/FunctorCategories.lean b/Mathlib/CategoryTheory/Idempotents/FunctorCategories.lean index 29b94beac78e3..34c4f44f5ce72 100644 --- a/Mathlib/CategoryTheory/Idempotents/FunctorCategories.lean +++ b/Mathlib/CategoryTheory/Idempotents/FunctorCategories.lean @@ -137,14 +137,11 @@ theorem toKaroubi_comp_karoubiFunctorCategoryEmbedding : apply Functor.ext · intro X Y f ext j - dsimp [toKaroubi] - simp only [eqToHom_app, eqToHom_refl] - erw [comp_id, id_comp] + simp · intro X apply Functor.ext · intro j j' φ ext - dsimp simp · intro j rfl diff --git a/Mathlib/CategoryTheory/Idempotents/FunctorExtension.lean b/Mathlib/CategoryTheory/Idempotents/FunctorExtension.lean index e227d08f5cbc7..8c6c94626e7a8 100644 --- a/Mathlib/CategoryTheory/Idempotents/FunctorExtension.lean +++ b/Mathlib/CategoryTheory/Idempotents/FunctorExtension.lean @@ -103,7 +103,7 @@ def functorExtension₁CompWhiskeringLeftToKaroubiIso : (fun X => { hom := { f := (F.obj X).p } inv := { f := (F.obj X).p } }) - (fun {X Y} f => by aesop_cat)) + (fun {X Y} f => by simp)) (by aesop_cat) /-- The counit isomorphism of the equivalence `(C ⥤ Karoubi D) ≌ (Karoubi C ⥤ Karoubi D)`. -/ @@ -180,7 +180,7 @@ def functorExtension₂CompWhiskeringLeftToKaroubiIso : (fun X => { hom := { f := 𝟙 _ } inv := { f := 𝟙 _ } }) - (by aesop_cat)) + (by simp)) (by aesop_cat) section IsIdempotentComplete diff --git a/Mathlib/CategoryTheory/Idempotents/HomologicalComplex.lean b/Mathlib/CategoryTheory/Idempotents/HomologicalComplex.lean index 52ef588b3f18e..3efddee1016cd 100644 --- a/Mathlib/CategoryTheory/Idempotents/HomologicalComplex.lean +++ b/Mathlib/CategoryTheory/Idempotents/HomologicalComplex.lean @@ -120,7 +120,7 @@ def inverse : HomologicalComplex (Karoubi C) c ⥤ Karoubi (HomologicalComplex C `Karoubi (HomologicalComplex C c) ≌ HomologicalComplex (Karoubi C) c`. -/ @[simps!] def counitIso : inverse ⋙ functor ≅ 𝟭 (HomologicalComplex (Karoubi C) c) := - eqToIso (Functor.ext (fun P => HomologicalComplex.ext (by aesop_cat) (by aesop_cat)) + eqToIso (Functor.ext (fun P => HomologicalComplex.ext (by aesop_cat) (by simp)) (by aesop_cat)) /-- The unit isomorphism of the equivalence diff --git a/Mathlib/CategoryTheory/Idempotents/Karoubi.lean b/Mathlib/CategoryTheory/Idempotents/Karoubi.lean index d4b4fc8a4f69e..aa8ebd6d96a7d 100644 --- a/Mathlib/CategoryTheory/Idempotents/Karoubi.lean +++ b/Mathlib/CategoryTheory/Idempotents/Karoubi.lean @@ -273,11 +273,11 @@ theorem decompId_p_toKaroubi (X : C) : decompId_p ((toKaroubi C).obj X) = 𝟙 _ theorem decompId_i_naturality {P Q : Karoubi C} (f : P ⟶ Q) : f ≫ decompId_i Q = decompId_i P ≫ (by exact Hom.mk f.f (by simp)) := by - aesop_cat + simp theorem decompId_p_naturality {P Q : Karoubi C} (f : P ⟶ Q) : decompId_p P ≫ f = (by exact Hom.mk f.f (by simp)) ≫ decompId_p Q := by - aesop_cat + simp @[simp] theorem zsmul_hom [Preadditive C] {P Q : Karoubi C} (n : ℤ) (f : P ⟶ Q) : (n • f).f = n • f.f := diff --git a/Mathlib/CategoryTheory/Idempotents/KaroubiKaroubi.lean b/Mathlib/CategoryTheory/Idempotents/KaroubiKaroubi.lean index 2fc3351f029a2..4195b5218d23f 100644 --- a/Mathlib/CategoryTheory/Idempotents/KaroubiKaroubi.lean +++ b/Mathlib/CategoryTheory/Idempotents/KaroubiKaroubi.lean @@ -46,7 +46,7 @@ instance [Preadditive C] : Functor.Additive (inverse C) where /-- The unit isomorphism of the equivalence -/ @[simps!] def unitIso : 𝟭 (Karoubi C) ≅ toKaroubi (Karoubi C) ⋙ inverse C := - eqToIso (Functor.ext (by aesop_cat) (by aesop_cat)) + eqToIso (Functor.ext (by aesop_cat) (by simp)) attribute [local simp] p_comm_f in /-- The counit isomorphism of the equivalence -/ diff --git a/Mathlib/CategoryTheory/IsConnected.lean b/Mathlib/CategoryTheory/IsConnected.lean index 73adc542b3de7..62109a61f38ae 100644 --- a/Mathlib/CategoryTheory/IsConnected.lean +++ b/Mathlib/CategoryTheory/IsConnected.lean @@ -102,7 +102,7 @@ def isoConstant [IsPreconnected J] {α : Type u₂} (F : J ⥤ Discrete α) (j : F ≅ (Functor.const J).obj (F.obj j) := (IsPreconnected.IsoConstantAux.factorThroughDiscrete F).symm ≪≫ isoWhiskerRight (IsPreconnected.iso_constant _ j).some _ - ≪≫ NatIso.ofComponents (fun _ => eqToIso Function.apply_invFun_apply) (by aesop_cat) + ≪≫ NatIso.ofComponents (fun _ => eqToIso Function.apply_invFun_apply) (by simp) /-- If `J` is connected, any functor to a discrete category is constant on objects. The converse is given in `IsConnected.of_any_functor_const_on_obj`. @@ -213,7 +213,7 @@ theorem isPreconnected_induction [IsPreconnected J] (Z : J → Sort*) {j₀ : J} (x : Z j₀) (j : J) : Nonempty (Z j) := (induct_on_objects { j | Nonempty (Z j) } ⟨x⟩ (fun f => ⟨by rintro ⟨y⟩; exact ⟨h₁ f y⟩, by rintro ⟨y⟩; exact ⟨h₂ f y⟩⟩) - j : _) + j :) /-- If `J` and `K` are equivalent, then if `J` is preconnected then `K` is as well. -/ theorem isPreconnected_of_equivalent {K : Type u₂} [Category.{v₂} K] [IsPreconnected J] @@ -373,7 +373,6 @@ theorem isPreconnected_zigzag [IsPreconnected J] (j₁ j₂ : J) : Zigzag j₁ j equiv_relation _ zigzag_equivalence (fun f => Relation.ReflTransGen.single (Or.inl (Nonempty.intro f))) _ _ -@[deprecated (since := "2024-02-19")] alias isConnected_zigzag := isPreconnected_zigzag theorem zigzag_isPreconnected (h : ∀ j₁ j₂ : J, Zigzag j₁ j₂) : IsPreconnected J := by apply IsPreconnected.of_constant_of_preserves_morphisms @@ -454,7 +453,4 @@ theorem nonempty_hom_of_preconnected_groupoid {G} [Groupoid G] [IsPreconnected G attribute [instance] nonempty_hom_of_preconnected_groupoid -@[deprecated (since := "2024-02-19")] -alias nonempty_hom_of_connected_groupoid := nonempty_hom_of_preconnected_groupoid - end CategoryTheory diff --git a/Mathlib/CategoryTheory/Limits/ColimitLimit.lean b/Mathlib/CategoryTheory/Limits/ColimitLimit.lean index 23386c1aed4c1..3622820f62ff3 100644 --- a/Mathlib/CategoryTheory/Limits/ColimitLimit.lean +++ b/Mathlib/CategoryTheory/Limits/ColimitLimit.lean @@ -104,7 +104,7 @@ noncomputable def colimitLimitToLimitColimitCone (G : J ⥤ K ⥤ C) [HasLimit G colim.mapCone (limit.cone G) ⟶ limit.cone (G ⋙ colim) where hom := colim.map (limitIsoSwapCompLim G).hom ≫ - colimitLimitToLimitColimit (uncurry.obj G : _) ≫ + colimitLimitToLimitColimit (uncurry.obj G :) ≫ lim.map (whiskerRight (currying.unitIso.app G).inv colim) w j := by dsimp diff --git a/Mathlib/CategoryTheory/Limits/Comma.lean b/Mathlib/CategoryTheory/Limits/Comma.lean index 93476394301e1..bef775a823d34 100644 --- a/Mathlib/CategoryTheory/Limits/Comma.lean +++ b/Mathlib/CategoryTheory/Limits/Comma.lean @@ -41,7 +41,7 @@ variable (F : J ⥤ Comma L R) in the comma category. -/ @[simps!] def limitAuxiliaryCone (c₁ : Cone (F ⋙ fst L R)) : Cone ((F ⋙ snd L R) ⋙ R) := - (Cones.postcompose (whiskerLeft F (Comma.natTrans L R) : _)).obj (L.mapCone c₁) + (Cones.postcompose (whiskerLeft F (Comma.natTrans L R) :)).obj (L.mapCone c₁) /-- If `R` preserves the appropriate limit, then given a cone for `F ⋙ fst L R : J ⥤ L` and a limit cone for `F ⋙ snd L R : J ⥤ R` we can build a cone for `F` which will turn out to be a limit @@ -87,7 +87,7 @@ noncomputable def coneOfPreservesIsLimit [PreservesLimit (F ⋙ snd L R) R] {c in the comma category. -/ @[simps!] def colimitAuxiliaryCocone (c₂ : Cocone (F ⋙ snd L R)) : Cocone ((F ⋙ fst L R) ⋙ L) := - (Cocones.precompose (whiskerLeft F (Comma.natTrans L R) : _)).obj (R.mapCocone c₂) + (Cocones.precompose (whiskerLeft F (Comma.natTrans L R) :)).obj (R.mapCocone c₂) /-- If `L` preserves the appropriate colimit, then given a colimit cocone for `F ⋙ fst L R : J ⥤ L` and @@ -210,7 +210,7 @@ noncomputable instance createsLimitsOfShape [PreservesLimitsOfShape J G] : CreatesLimitsOfShape J (proj X G) where noncomputable instance createsLimitsOfSize [PreservesLimitsOfSize.{w, w'} G] : - CreatesLimitsOfSize.{w, w'} (proj X G : _) where + CreatesLimitsOfSize.{w, w'} (proj X G :) where instance mono_right_of_mono [HasPullbacks A] [PreservesLimitsOfShape WalkingCospan G] {Y Z : StructuredArrow X G} (f : Y ⟶ Z) [Mono f] : Mono f.right := @@ -255,7 +255,7 @@ noncomputable instance createsColimitsOfShape [PreservesColimitsOfShape J G] : CreatesColimitsOfShape J (proj G X) where noncomputable instance createsColimitsOfSize [PreservesColimitsOfSize.{w, w'} G] : - CreatesColimitsOfSize.{w, w'} (proj G X : _) where + CreatesColimitsOfSize.{w, w'} (proj G X :) where instance epi_left_of_epi [HasPushouts A] [PreservesColimitsOfShape WalkingSpan G] {Y Z : CostructuredArrow G X} (f : Y ⟶ Z) [Epi f] : Epi f.left := diff --git a/Mathlib/CategoryTheory/Limits/Cones.lean b/Mathlib/CategoryTheory/Limits/Cones.lean index 6fee46c53e64c..05e9c7a87a567 100644 --- a/Mathlib/CategoryTheory/Limits/Cones.lean +++ b/Mathlib/CategoryTheory/Limits/Cones.lean @@ -158,7 +158,7 @@ instance inhabitedCocone (F : Discrete PUnit ⥤ C) : Inhabited (Cocone F) := intro X Y f match X, Y, f with | .mk A, .mk B, .up g => - aesop_cat + simp } }⟩ @@ -408,7 +408,7 @@ def functoriality : Cone F ⥤ Cone (F ⋙ G) where { pt := G.obj A.pt π := { app := fun j => G.map (A.π.app j) - naturality := by intros; erw [← G.map_comp]; aesop_cat } } + naturality := by intros; erw [← G.map_comp]; simp } } map f := { hom := G.map f.hom w := fun j => by simp [-ConeMorphism.w, ← f.w j] } @@ -606,7 +606,7 @@ def functoriality : Cocone F ⥤ Cocone (F ⋙ G) where { pt := G.obj A.pt ι := { app := fun j => G.map (A.ι.app j) - naturality := by intros; erw [← G.map_comp]; aesop_cat } } + naturality := by intros; erw [← G.map_comp]; simp } } map f := { hom := G.map f.hom w := by intros; rw [← Functor.map_comp, CoconeMorphism.w] } @@ -719,7 +719,7 @@ isomorphic to the cone `H'.mapCone`. -/ @[simps!] def postcomposeWhiskerLeftMapCone {H H' : C ⥤ D} (α : H ≅ H') (c : Cone F) : - (Cones.postcompose (whiskerLeft F α.hom : _)).obj (mapCone H c) ≅ mapCone H' c := + (Cones.postcompose (whiskerLeft F α.hom :)).obj (mapCone H c) ≅ mapCone H' c := (functorialityCompPostcompose α).app c /-- @@ -730,7 +730,7 @@ a cone over `G ⋙ H`, and they are both isomorphic. @[simps!] def mapConePostcompose {α : F ⟶ G} {c} : mapCone H ((Cones.postcompose α).obj c) ≅ - (Cones.postcompose (whiskerRight α H : _)).obj (mapCone H c) := + (Cones.postcompose (whiskerRight α H :)).obj (mapCone H c) := Cones.ext (Iso.refl _) /-- `mapCone` commutes with `postcomposeEquivalence` @@ -738,7 +738,7 @@ def mapConePostcompose {α : F ⟶ G} {c} : @[simps!] def mapConePostcomposeEquivalenceFunctor {α : F ≅ G} {c} : mapCone H ((Cones.postcomposeEquivalence α).functor.obj c) ≅ - (Cones.postcomposeEquivalence (isoWhiskerRight α H : _)).functor.obj (mapCone H c) := + (Cones.postcomposeEquivalence (isoWhiskerRight α H :)).functor.obj (mapCone H c) := Cones.ext (Iso.refl _) /-- `functoriality F _ ⋙ precompose (whiskerLeft F _)` simplifies to `functoriality F _`. -/ @@ -755,7 +755,7 @@ isomorphic to the cocone `H'.mapCocone`. -/ @[simps!] def precomposeWhiskerLeftMapCocone {H H' : C ⥤ D} (α : H ≅ H') (c : Cocone F) : - (Cocones.precompose (whiskerLeft F α.inv : _)).obj (mapCocone H c) ≅ mapCocone H' c := + (Cocones.precompose (whiskerLeft F α.inv :)).obj (mapCocone H c) ≅ mapCocone H' c := (functorialityCompPrecompose α).app c /-- `map_cocone` commutes with `precompose`. In particular, for `F : J ⥤ C`, given a cocone @@ -765,7 +765,7 @@ ways of producing a cocone over `G ⋙ H`, and they are both isomorphic. @[simps!] def mapCoconePrecompose {α : F ⟶ G} {c} : mapCocone H ((Cocones.precompose α).obj c) ≅ - (Cocones.precompose (whiskerRight α H : _)).obj (mapCocone H c) := + (Cocones.precompose (whiskerRight α H :)).obj (mapCocone H c) := Cocones.ext (Iso.refl _) /-- `mapCocone` commutes with `precomposeEquivalence` @@ -773,7 +773,7 @@ def mapCoconePrecompose {α : F ⟶ G} {c} : @[simps!] def mapCoconePrecomposeEquivalenceFunctor {α : F ≅ G} {c} : mapCocone H ((Cocones.precomposeEquivalence α).functor.obj c) ≅ - (Cocones.precomposeEquivalence (isoWhiskerRight α H : _)).functor.obj (mapCocone H c) := + (Cocones.precomposeEquivalence (isoWhiskerRight α H :)).functor.obj (mapCocone H c) := Cocones.ext (Iso.refl _) /-- `mapCone` commutes with `whisker` diff --git a/Mathlib/CategoryTheory/Limits/Constructions/BinaryProducts.lean b/Mathlib/CategoryTheory/Limits/Constructions/BinaryProducts.lean index 4ad7a58d9ac61..f0c238e0c8044 100644 --- a/Mathlib/CategoryTheory/Limits/Constructions/BinaryProducts.lean +++ b/Mathlib/CategoryTheory/Limits/Constructions/BinaryProducts.lean @@ -28,7 +28,7 @@ variable {C : Type u} [Category.{v} C] {D : Type u'} [Category.{v'} D] (F : C def isBinaryProductOfIsTerminalIsPullback (F : Discrete WalkingPair ⥤ C) (c : Cone F) {X : C} (hX : IsTerminal X) (f : F.obj ⟨WalkingPair.left⟩ ⟶ X) (g : F.obj ⟨WalkingPair.right⟩ ⟶ X) (hc : IsLimit - (PullbackCone.mk (c.π.app ⟨WalkingPair.left⟩) (c.π.app ⟨WalkingPair.right⟩ : _) <| + (PullbackCone.mk (c.π.app ⟨WalkingPair.left⟩) (c.π.app ⟨WalkingPair.right⟩ :) <| hX.hom_ext (_ ≫ f) (_ ≫ g))) : IsLimit c where lift s := hc.lift @@ -38,7 +38,7 @@ def isBinaryProductOfIsTerminalIsPullback (F : Discrete WalkingPair ⥤ C) (c : WalkingPair.casesOn j (hc.fac _ WalkingCospan.left) (hc.fac _ WalkingCospan.right) uniq s m J := by let c' := - PullbackCone.mk (m ≫ c.π.app ⟨WalkingPair.left⟩) (m ≫ c.π.app ⟨WalkingPair.right⟩ : _) + PullbackCone.mk (m ≫ c.π.app ⟨WalkingPair.left⟩) (m ≫ c.π.app ⟨WalkingPair.right⟩ :) (hX.hom_ext (_ ≫ f) (_ ≫ g)) dsimp; rw [← J, ← J] apply hc.hom_ext @@ -134,7 +134,7 @@ def isBinaryCoproductOfIsInitialIsPushout (F : Discrete WalkingPair ⥤ C) (c : (hX : IsInitial X) (f : X ⟶ F.obj ⟨WalkingPair.left⟩) (g : X ⟶ F.obj ⟨WalkingPair.right⟩) (hc : IsColimit - (PushoutCocone.mk (c.ι.app ⟨WalkingPair.left⟩) (c.ι.app ⟨WalkingPair.right⟩ : _) <| + (PushoutCocone.mk (c.ι.app ⟨WalkingPair.left⟩) (c.ι.app ⟨WalkingPair.right⟩ :) <| hX.hom_ext (f ≫ _) (g ≫ _))) : IsColimit c where desc s := diff --git a/Mathlib/CategoryTheory/Limits/Constructions/Filtered.lean b/Mathlib/CategoryTheory/Limits/Constructions/Filtered.lean index 5712931af86a2..12841a497aa58 100644 --- a/Mathlib/CategoryTheory/Limits/Constructions/Filtered.lean +++ b/Mathlib/CategoryTheory/Limits/Constructions/Filtered.lean @@ -68,7 +68,7 @@ def liftToFinsetColimitCocone [HasColimitsOfShape (Finset (Discrete α)) C] convert h j using 1 · simp [← colimit.w (liftToFinsetObj F) ⟨⟨Finset.singleton_subset_iff.2 hj⟩⟩] rfl - · aesop_cat } + · simp } variable (C) (α) @@ -180,7 +180,7 @@ def liftToFinsetLimitCone [HasLimitsOfShape (Finset (Discrete α))ᵒᵖ C] convert h j using 1 · simp [← limit.w (liftToFinsetObj F) ⟨⟨⟨Finset.singleton_subset_iff.2 hj⟩⟩⟩] rfl - · aesop_cat } + · simp } variable (C) (α) diff --git a/Mathlib/CategoryTheory/Limits/Constructions/ZeroObjects.lean b/Mathlib/CategoryTheory/Limits/Constructions/ZeroObjects.lean index d978a5c54dda2..e9e831e516ae5 100644 --- a/Mathlib/CategoryTheory/Limits/Constructions/ZeroObjects.lean +++ b/Mathlib/CategoryTheory/Limits/Constructions/ZeroObjects.lean @@ -32,7 +32,7 @@ def binaryFanZeroLeft (X : C) : BinaryFan (0 : C) X := /-- The limit cone for the product with a zero object is limiting. -/ def binaryFanZeroLeftIsLimit (X : C) : IsLimit (binaryFanZeroLeft X) := - BinaryFan.isLimitMk (fun s => BinaryFan.snd s) (by aesop_cat) (by aesop_cat) + BinaryFan.isLimitMk (fun s => BinaryFan.snd s) (by aesop_cat) (by simp) (fun s m _ h₂ => by simpa using h₂) instance hasBinaryProduct_zero_left (X : C) : HasBinaryProduct (0 : C) X := @@ -57,7 +57,7 @@ def binaryFanZeroRight (X : C) : BinaryFan X (0 : C) := /-- The limit cone for the product with a zero object is limiting. -/ def binaryFanZeroRightIsLimit (X : C) : IsLimit (binaryFanZeroRight X) := - BinaryFan.isLimitMk (fun s => BinaryFan.fst s) (by aesop_cat) (by aesop_cat) + BinaryFan.isLimitMk (fun s => BinaryFan.fst s) (by simp) (by aesop_cat) (fun s m h₁ _ => by simpa using h₁) instance hasBinaryProduct_zero_right (X : C) : HasBinaryProduct X (0 : C) := @@ -82,7 +82,7 @@ def binaryCofanZeroLeft (X : C) : BinaryCofan (0 : C) X := /-- The colimit cocone for the coproduct with a zero object is colimiting. -/ def binaryCofanZeroLeftIsColimit (X : C) : IsColimit (binaryCofanZeroLeft X) := - BinaryCofan.isColimitMk (fun s => BinaryCofan.inr s) (by aesop_cat) (by aesop_cat) + BinaryCofan.isColimitMk (fun s => BinaryCofan.inr s) (by aesop_cat) (by simp) (fun s m _ h₂ => by simpa using h₂) instance hasBinaryCoproduct_zero_left (X : C) : HasBinaryCoproduct (0 : C) X := @@ -107,7 +107,7 @@ def binaryCofanZeroRight (X : C) : BinaryCofan X (0 : C) := /-- The colimit cocone for the coproduct with a zero object is colimiting. -/ def binaryCofanZeroRightIsColimit (X : C) : IsColimit (binaryCofanZeroRight X) := - BinaryCofan.isColimitMk (fun s => BinaryCofan.inl s) (by aesop_cat) (by aesop_cat) + BinaryCofan.isColimitMk (fun s => BinaryCofan.inl s) (by simp) (by aesop_cat) (fun s m h₁ _ => by simpa using h₁) instance hasBinaryCoproduct_zero_right (X : C) : HasBinaryCoproduct X (0 : C) := diff --git a/Mathlib/CategoryTheory/Limits/Creates.lean b/Mathlib/CategoryTheory/Limits/Creates.lean index 9769ea2520000..5f74ad2e03a1c 100644 --- a/Mathlib/CategoryTheory/Limits/Creates.lean +++ b/Mathlib/CategoryTheory/Limits/Creates.lean @@ -500,7 +500,7 @@ def createsLimitOfIsoDiagram {K₁ K₂ : J ⥤ C} (F : C ⥤ D) (h : K₁ ≅ K CreatesLimit K₂ F := { reflectsLimit_of_iso_diagram F h with lifts := fun c t => - let t' := (IsLimit.postcomposeInvEquiv (isoWhiskerRight h F : _) c).symm t + let t' := (IsLimit.postcomposeInvEquiv (isoWhiskerRight h F :) c).symm t { liftedCone := (Cones.postcompose h.hom).obj (liftLimit t') validLift := Functor.mapConePostcompose F ≪≫ @@ -513,7 +513,7 @@ def createsLimitOfIsoDiagram {K₁ K₂ : J ⥤ C} (F : C ⥤ D) (h : K₁ ≅ K /-- If `F` creates the limit of `K` and `F ≅ G`, then `G` creates the limit of `K`. -/ def createsLimitOfNatIso {F G : C ⥤ D} (h : F ≅ G) [CreatesLimit K F] : CreatesLimit K G where lifts c t := - { liftedCone := liftLimit ((IsLimit.postcomposeInvEquiv (isoWhiskerLeft K h : _) c).symm t) + { liftedCone := liftLimit ((IsLimit.postcomposeInvEquiv (isoWhiskerLeft K h :) c).symm t) validLift := by refine (IsLimit.mapConeEquiv h ?_).uniqueUpToIso t apply IsLimit.ofIsoLimit _ (liftedLimitMapsToOriginal _).symm @@ -534,7 +534,7 @@ def createsColimitOfIsoDiagram {K₁ K₂ : J ⥤ C} (F : C ⥤ D) (h : K₁ ≅ CreatesColimit K₂ F := { reflectsColimit_of_iso_diagram F h with lifts := fun c t => - let t' := (IsColimit.precomposeHomEquiv (isoWhiskerRight h F : _) c).symm t + let t' := (IsColimit.precomposeHomEquiv (isoWhiskerRight h F :) c).symm t { liftedCocone := (Cocones.precompose h.inv).obj (liftColimit t') validLift := Functor.mapCoconePrecompose F ≪≫ @@ -548,7 +548,7 @@ def createsColimitOfIsoDiagram {K₁ K₂ : J ⥤ C} (F : C ⥤ D) (h : K₁ ≅ /-- If `F` creates the colimit of `K` and `F ≅ G`, then `G` creates the colimit of `K`. -/ def createsColimitOfNatIso {F G : C ⥤ D} (h : F ≅ G) [CreatesColimit K F] : CreatesColimit K G where lifts c t := - { liftedCocone := liftColimit ((IsColimit.precomposeHomEquiv (isoWhiskerLeft K h : _) c).symm t) + { liftedCocone := liftColimit ((IsColimit.precomposeHomEquiv (isoWhiskerLeft K h :) c).symm t) validLift := by refine (IsColimit.mapCoconeEquiv h ?_).uniqueUpToIso t apply IsColimit.ofIsoColimit _ (liftedColimitMapsToOriginal _).symm diff --git a/Mathlib/CategoryTheory/Limits/Final.lean b/Mathlib/CategoryTheory/Limits/Final.lean index f704225f03972..5124dc554a97f 100644 --- a/Mathlib/CategoryTheory/Limits/Final.lean +++ b/Mathlib/CategoryTheory/Limits/Final.lean @@ -257,7 +257,7 @@ theorem colimit_cocone_comp_aux (s : Cocone (F ⋙ G)) (j : C) : G.map (homToLift F (F.obj j)) ≫ s.ι.app (lift F (F.obj j)) = s.ι.app j := by -- This point is that this would be true if we took `lift (F.obj j)` to just be `j` -- and `homToLift (F.obj j)` to be `𝟙 (F.obj j)`. - apply induction F fun X k => G.map k ≫ s.ι.app X = (s.ι.app j : _) + apply induction F fun X k => G.map k ≫ s.ι.app X = (s.ι.app j :) · intro j₁ j₂ k₁ k₂ f w h rw [← w] rw [← s.w f] at h @@ -607,7 +607,7 @@ theorem limit_cone_comp_aux (s : Cone (F ⋙ G)) (j : C) : s.π.app (lift F (F.obj j)) ≫ G.map (homToLift F (F.obj j)) = s.π.app j := by -- This point is that this would be true if we took `lift (F.obj j)` to just be `j` -- and `homToLift (F.obj j)` to be `𝟙 (F.obj j)`. - apply induction F fun X k => s.π.app X ≫ G.map k = (s.π.app j : _) + apply induction F fun X k => s.π.app X ≫ G.map k = (s.π.app j :) · intro j₁ j₂ k₁ k₂ f w h rw [← s.w f] rw [← w] at h diff --git a/Mathlib/CategoryTheory/Limits/Fubini.lean b/Mathlib/CategoryTheory/Limits/Fubini.lean index 898d073faa5a4..9a6aedbc78108 100644 --- a/Mathlib/CategoryTheory/Limits/Fubini.lean +++ b/Mathlib/CategoryTheory/Limits/Fubini.lean @@ -35,14 +35,12 @@ All statements have their counterpart for colimits. -/ -universe v u - open CategoryTheory namespace CategoryTheory.Limits -variable {J K : Type v} [SmallCategory J] [SmallCategory K] -variable {C : Type u} [Category.{v} C] +variable {J K : Type*} [Category J] [Category K] +variable {C : Type*} [Category C] variable (F : J ⥤ K ⥤ C) -- We could try introducing a "dependent functor type" to handle this? @@ -486,9 +484,9 @@ end section -variable [HasLimits C] +variable [HasLimitsOfShape K C] [HasLimitsOfShape J C] [HasLimitsOfShape (K × J) C] [HasLimit G] + [HasLimit (curry.obj G ⋙ lim)] --- Certainly one could weaken the hypotheses here. open CategoryTheory.prod /-- A variant of the Fubini theorem for a functor `G : J × K ⥤ C`, @@ -525,7 +523,8 @@ end section -variable [HasColimits C] +variable [HasColimitsOfShape K C] [HasColimitsOfShape J C] [HasColimitsOfShape (K × J) C] + [HasColimit G] [HasColimit (curry.obj G ⋙ colim)] open CategoryTheory.prod diff --git a/Mathlib/CategoryTheory/Limits/FullSubcategory.lean b/Mathlib/CategoryTheory/Limits/FullSubcategory.lean index 18abbde52961c..11c03c8e182d8 100644 --- a/Mathlib/CategoryTheory/Limits/FullSubcategory.lean +++ b/Mathlib/CategoryTheory/Limits/FullSubcategory.lean @@ -113,8 +113,6 @@ theorem hasLimit_of_closedUnderLimits (h : ClosedUnderLimitsOfShape J P) have : CreatesLimit F (fullSubcategoryInclusion P) := createsLimitFullSubcategoryInclusionOfClosed h F hasLimit_of_created F (fullSubcategoryInclusion P) -@[deprecated (since := "2024-03-23")] -alias hasLimit_of_closed_under_limits := hasLimit_of_closedUnderLimits theorem hasLimitsOfShape_of_closedUnderLimits (h : ClosedUnderLimitsOfShape J P) [HasLimitsOfShape J C] : HasLimitsOfShape J (FullSubcategory P) := @@ -136,14 +134,10 @@ theorem hasColimit_of_closedUnderColimits (h : ClosedUnderColimitsOfShape J P) have : CreatesColimit F (fullSubcategoryInclusion P) := createsColimitFullSubcategoryInclusionOfClosed h F hasColimit_of_created F (fullSubcategoryInclusion P) -@[deprecated (since := "2024-03-23")] -alias hasColimit_of_closed_under_colimits := hasColimit_of_closedUnderColimits theorem hasColimitsOfShape_of_closedUnderColimits (h : ClosedUnderColimitsOfShape J P) [HasColimitsOfShape J C] : HasColimitsOfShape J (FullSubcategory P) := { has_colimit := fun F => hasColimit_of_closedUnderColimits h F } -@[deprecated (since := "2024-03-23")] -alias hasColimitsOfShape_of_closed_under_colimits := hasColimitsOfShape_of_closedUnderColimits end diff --git a/Mathlib/CategoryTheory/Limits/HasLimits.lean b/Mathlib/CategoryTheory/Limits/HasLimits.lean index 27bcb80063f88..e9ece3954362f 100644 --- a/Mathlib/CategoryTheory/Limits/HasLimits.lean +++ b/Mathlib/CategoryTheory/Limits/HasLimits.lean @@ -225,13 +225,13 @@ def limit.isoLimitCone {F : J ⥤ C} [HasLimit F] (t : LimitCone F) : limit F theorem limit.isoLimitCone_hom_π {F : J ⥤ C} [HasLimit F] (t : LimitCone F) (j : J) : (limit.isoLimitCone t).hom ≫ t.cone.π.app j = limit.π F j := by dsimp [limit.isoLimitCone, IsLimit.conePointUniqueUpToIso] - aesop_cat + simp @[reassoc (attr := simp)] theorem limit.isoLimitCone_inv_π {F : J ⥤ C} [HasLimit F] (t : LimitCone F) (j : J) : (limit.isoLimitCone t).inv ≫ limit.π F j = t.cone.π.app j := by dsimp [limit.isoLimitCone, IsLimit.conePointUniqueUpToIso] - aesop_cat + simp @[ext] theorem limit.hom_ext {F : J ⥤ C} [HasLimit F] {X : C} {f f' : X ⟶ limit F} @@ -458,10 +458,10 @@ def lim : (J ⥤ C) ⥤ C where map α := limMap α map_id F := by apply Limits.limit.hom_ext; intro j - erw [limMap_π, Category.id_comp, Category.comp_id] + simp map_comp α β := by apply Limits.limit.hom_ext; intro j - erw [assoc, IsLimit.fac, IsLimit.fac, ← assoc, IsLimit.fac, assoc]; rfl + simp [assoc] end @@ -750,13 +750,13 @@ def colimit.isoColimitCocone {F : J ⥤ C} [HasColimit F] (t : ColimitCocone F) theorem colimit.isoColimitCocone_ι_hom {F : J ⥤ C} [HasColimit F] (t : ColimitCocone F) (j : J) : colimit.ι F j ≫ (colimit.isoColimitCocone t).hom = t.cocone.ι.app j := by dsimp [colimit.isoColimitCocone, IsColimit.coconePointUniqueUpToIso] - aesop_cat + simp @[reassoc (attr := simp)] theorem colimit.isoColimitCocone_ι_inv {F : J ⥤ C} [HasColimit F] (t : ColimitCocone F) (j : J) : t.cocone.ι.app j ≫ (colimit.isoColimitCocone t).inv = colimit.ι F j := by dsimp [colimit.isoColimitCocone, IsColimit.coconePointUniqueUpToIso] - aesop_cat + simp @[ext] theorem colimit.hom_ext {F : J ⥤ C} [HasColimit F] {X : C} {f f' : colimit F ⟶ X} diff --git a/Mathlib/CategoryTheory/Limits/Indization/Category.lean b/Mathlib/CategoryTheory/Limits/Indization/Category.lean index c927eac12564d..4de4c45fb1b0f 100644 --- a/Mathlib/CategoryTheory/Limits/Indization/Category.lean +++ b/Mathlib/CategoryTheory/Limits/Indization/Category.lean @@ -8,6 +8,7 @@ import Mathlib.CategoryTheory.Limits.Constructions.Filtered import Mathlib.CategoryTheory.Limits.FullSubcategory import Mathlib.CategoryTheory.Limits.Indization.LocallySmall import Mathlib.CategoryTheory.Limits.Indization.FilteredColimits +import Mathlib.CategoryTheory.Limits.Indization.Products /-! # The category of Ind-objects @@ -24,7 +25,10 @@ Adopting the theorem numbering of [Kashiwara2006], we show that * `C ⥤ Ind C` preserves finite colimits (6.1.6), * if `C` has coproducts indexed by a finite type `α`, then `Ind C` has coproducts indexed by `α` (6.1.18(ii)), -* if `C` has finite coproducts, then `Ind C` has small coproducts (6.1.18(ii)). +* if `C` has finite coproducts, then `Ind C` has small coproducts (6.1.18(ii)), +* if `C` has products indexed by `α`, then `Ind C` has products indexed by `α`, and the functor + `Ind C ⥤ Cᵒᵖ ⥤ Type v` creates such products (6.1.17) +* the functor `C ⥤ Ind C` preserves small limits. More limit-colimit properties will follow. @@ -107,6 +111,23 @@ instance : HasFilteredColimits (Ind C) where HasColimitsOfShape _ _ _ := hasColimitsOfShape_of_hasColimitsOfShape_createsColimitsOfShape (Ind.inclusion C) +noncomputable instance {J : Type v} [HasLimitsOfShape (Discrete J) C] : + CreatesLimitsOfShape (Discrete J) (Ind.inclusion C) := + letI _ : CreatesLimitsOfShape (Discrete J) (fullSubcategoryInclusion (IsIndObject (C := C))) := + createsLimitsOfShapeFullSubcategoryInclusion (closedUnderLimitsOfShape_of_limit + (isIndObject_limit_of_discrete_of_hasLimitsOfShape _)) + inferInstanceAs <| + CreatesLimitsOfShape (Discrete J) ((Ind.equivalence C).functor ⋙ fullSubcategoryInclusion _) + +instance {J : Type v} [HasLimitsOfShape (Discrete J) C] : + HasLimitsOfShape (Discrete J) (Ind C) := + hasLimitsOfShape_of_hasLimitsOfShape_createsLimitsOfShape (Ind.inclusion C) + +instance : PreservesLimits (Ind.yoneda (C := C)) := + letI _ : PreservesLimitsOfSize.{v, v} (Ind.yoneda ⋙ Ind.inclusion C) := + preservesLimits_of_natIso Ind.yonedaCompInclusion.symm + preservesLimits_of_reflects_of_preserves Ind.yoneda (Ind.inclusion C) + theorem Ind.isIndObject_inclusion_obj (X : Ind C) : IsIndObject ((Ind.inclusion C).obj X) := X.2 diff --git a/Mathlib/CategoryTheory/Limits/IsLimit.lean b/Mathlib/CategoryTheory/Limits/IsLimit.lean index d0aebd2e0120a..a15e2af7ecdaa 100644 --- a/Mathlib/CategoryTheory/Limits/IsLimit.lean +++ b/Mathlib/CategoryTheory/Limits/IsLimit.lean @@ -263,8 +263,8 @@ def conePointsIsoOfNatIso {F G : J ⥤ C} {s : Cone F} {t : Cone G} (P : IsLimit (w : F ≅ G) : s.pt ≅ t.pt where hom := Q.map s w.hom inv := P.map t w.inv - hom_inv_id := P.hom_ext (by aesop_cat) - inv_hom_id := Q.hom_ext (by aesop_cat) + hom_inv_id := P.hom_ext (by simp) + inv_hom_id := Q.hom_ext (by simp) @[reassoc] theorem conePointsIsoOfNatIso_hom_comp {F G : J ⥤ C} {s : Cone F} {t : Cone G} (P : IsLimit s) @@ -404,7 +404,7 @@ def ofFaithful {t : Cone F} {D : Type u₄} [Category.{v₄} D] (G : C ⥤ D) [G -/ def mapConeEquiv {D : Type u₄} [Category.{v₄} D] {K : J ⥤ C} {F G : C ⥤ D} (h : F ≅ G) {c : Cone K} (t : IsLimit (mapCone F c)) : IsLimit (mapCone G c) := by - apply postcomposeInvEquiv (isoWhiskerLeft K h : _) (mapCone G c) _ + apply postcomposeInvEquiv (isoWhiskerLeft K h :) (mapCone G c) _ apply t.ofIsoLimit (postcomposeWhiskerLeftMapCone h.symm c).symm /-- A cone is a limit cone exactly if @@ -442,7 +442,7 @@ theorem coneOfHom_homOfCone (s : Cone F) : coneOfHom h (homOfCone h s) = s := by @[simp] theorem homOfCone_coneOfHom {Y : C} (f : Y ⟶ X) : homOfCone h (coneOfHom h f) = f := - congrArg ULift.down (congrFun (congrFun (congrArg NatTrans.app h.hom_inv_id) (op Y)) ⟨f⟩ : _) + congrArg ULift.down (congrFun (congrFun (congrArg NatTrans.app h.hom_inv_id) (op Y)) ⟨f⟩ :) /-- If `F.cones` is represented by `X`, the cone corresponding to the identity morphism on `X` will be a limit cone. -/ @@ -643,14 +643,10 @@ theorem hom_desc (h : IsColimit t) {W : C} (m : t.pt ⟶ W) : m = h.desc { pt := W - ι := - { app := fun b => t.ι.app b ≫ m - naturality := by intros; erw [← assoc, t.ι.naturality, comp_id, comp_id] } } := + ι := { app := fun b => t.ι.app b ≫ m } } := h.uniq { pt := W - ι := - { app := fun b => t.ι.app b ≫ m - naturality := _ } } + ι := { app := fun b => t.ι.app b ≫ m } } m fun _ => rfl /-- Two morphisms out of a colimit are equal if their compositions with @@ -723,8 +719,8 @@ def coconePointsIsoOfNatIso {F G : J ⥤ C} {s : Cocone F} {t : Cocone G} (P : I (Q : IsColimit t) (w : F ≅ G) : s.pt ≅ t.pt where hom := P.map t w.hom inv := Q.map s w.inv - hom_inv_id := P.hom_ext (by aesop_cat) - inv_hom_id := Q.hom_ext (by aesop_cat) + hom_inv_id := P.hom_ext (by simp) + inv_hom_id := Q.hom_ext (by simp) @[reassoc] theorem comp_coconePointsIsoOfNatIso_hom {F G : J ⥤ C} {s : Cocone F} {t : Cocone G} @@ -876,7 +872,7 @@ def ofFaithful {t : Cocone F} {D : Type u₄} [Category.{v₄} D] (G : C ⥤ D) def mapCoconeEquiv {D : Type u₄} [Category.{v₄} D] {K : J ⥤ C} {F G : C ⥤ D} (h : F ≅ G) {c : Cocone K} (t : IsColimit (mapCocone F c)) : IsColimit (mapCocone G c) := by apply IsColimit.ofIsoColimit _ (precomposeWhiskerLeftMapCocone h c) - apply (precomposeInvEquiv (isoWhiskerLeft K h : _) _).symm t + apply (precomposeInvEquiv (isoWhiskerLeft K h :) _).symm t /-- A cocone is a colimit cocone exactly if there is a unique cocone morphism from any other cocone. @@ -912,7 +908,7 @@ theorem coconeOfHom_homOfCocone (s : Cocone F) : coconeOfHom h (homOfCocone h s) @[simp] theorem homOfCocone_cooneOfHom {Y : C} (f : X ⟶ Y) : homOfCocone h (coconeOfHom h f) = f := - congrArg ULift.down (congrFun (congrFun (congrArg NatTrans.app h.hom_inv_id) Y) ⟨f⟩ : _) + congrArg ULift.down (congrFun (congrFun (congrArg NatTrans.app h.hom_inv_id) Y) ⟨f⟩ :) /-- If `F.cocones` is corepresented by `X`, the cocone corresponding to the identity morphism on `X` will be a colimit cocone. -/ diff --git a/Mathlib/CategoryTheory/Limits/MonoCoprod.lean b/Mathlib/CategoryTheory/Limits/MonoCoprod.lean index 9922662f9f26a..868732f7d96d1 100644 --- a/Mathlib/CategoryTheory/Limits/MonoCoprod.lean +++ b/Mathlib/CategoryTheory/Limits/MonoCoprod.lean @@ -60,7 +60,7 @@ theorem binaryCofan_inr {A B : C} [MonoCoprod C] (c : BinaryCofan A B) (hc : IsC Mono c.inr := by haveI hc' : IsColimit (BinaryCofan.mk c.inr c.inl) := BinaryCofan.IsColimit.mk _ (fun f₁ f₂ => hc.desc (BinaryCofan.mk f₂ f₁)) - (by aesop_cat) (by aesop_cat) + (by simp) (by simp) (fun f₁ f₂ m h₁ h₂ => BinaryCofan.IsColimit.hom_ext hc (by aesop_cat) (by aesop_cat)) exact binaryCofan_inl _ hc' diff --git a/Mathlib/CategoryTheory/Limits/Opposites.lean b/Mathlib/CategoryTheory/Limits/Opposites.lean index 6c32f2d694aba..7c040f35246ea 100644 --- a/Mathlib/CategoryTheory/Limits/Opposites.lean +++ b/Mathlib/CategoryTheory/Limits/Opposites.lean @@ -32,10 +32,6 @@ namespace CategoryTheory.Limits variable {C : Type u₁} [Category.{v₁} C] variable {J : Type u₂} [Category.{v₂} J] -@[deprecated (since := "2024-03-26")] alias isLimitCoconeOp := IsColimit.op -@[deprecated (since := "2024-03-26")] alias isColimitConeOp := IsLimit.op -@[deprecated (since := "2024-03-26")] alias isLimitCoconeUnop := IsColimit.unop -@[deprecated (since := "2024-03-26")] alias isColimitConeUnop := IsLimit.unop /-- Turn a colimit for `F : J ⥤ Cᵒᵖ` into a limit for `F.leftOp : Jᵒᵖ ⥤ C`. -/ @[simps] @@ -888,7 +884,7 @@ theorem unop_fst {X Y Z : Cᵒᵖ} {f : X ⟶ Y} {g : X ⟶ Z} (c : PushoutCocon c.unop.fst = c.inl.unop := by simp theorem unop_snd {X Y Z : Cᵒᵖ} {f : X ⟶ Y} {g : X ⟶ Z} (c : PushoutCocone f g) : - c.unop.snd = c.inr.unop := by aesop_cat + c.unop.snd = c.inr.unop := by simp -- Porting note: it was originally @[simps (config := lemmasOnly)] /-- The obvious map `PushoutCocone f.op g.op → PullbackCone f g` -/ @@ -898,10 +894,10 @@ def op {X Y Z : C} {f : X ⟶ Y} {g : X ⟶ Z} (c : PushoutCocone f g) : Pullbac (Cone.whisker walkingSpanOpEquiv.inverse (Cocone.op c)) theorem op_fst {X Y Z : C} {f : X ⟶ Y} {g : X ⟶ Z} (c : PushoutCocone f g) : - c.op.fst = c.inl.op := by aesop_cat + c.op.fst = c.inl.op := by simp theorem op_snd {X Y Z : C} {f : X ⟶ Y} {g : X ⟶ Z} (c : PushoutCocone f g) : - c.op.snd = c.inr.op := by aesop_cat + c.op.snd = c.inr.op := by simp end PushoutCocone @@ -917,10 +913,10 @@ def unop {X Y Z : Cᵒᵖ} {f : X ⟶ Z} {g : Y ⟶ Z} (c : PullbackCone f g) : (Cone.whisker walkingSpanOpEquiv.functor c)) theorem unop_inl {X Y Z : Cᵒᵖ} {f : X ⟶ Z} {g : Y ⟶ Z} (c : PullbackCone f g) : - c.unop.inl = c.fst.unop := by aesop_cat + c.unop.inl = c.fst.unop := by simp theorem unop_inr {X Y Z : Cᵒᵖ} {f : X ⟶ Z} {g : Y ⟶ Z} (c : PullbackCone f g) : - c.unop.inr = c.snd.unop := by aesop_cat + c.unop.inr = c.snd.unop := by simp /-- The obvious map `PullbackCone f g → PushoutCocone f.op g.op` -/ @[simps!] @@ -929,10 +925,10 @@ def op {X Y Z : C} {f : X ⟶ Z} {g : Y ⟶ Z} (c : PullbackCone f g) : PushoutC (Cocone.whisker walkingCospanOpEquiv.inverse (Cone.op c)) theorem op_inl {X Y Z : C} {f : X ⟶ Z} {g : Y ⟶ Z} (c : PullbackCone f g) : - c.op.inl = c.fst.op := by aesop_cat + c.op.inl = c.fst.op := by simp theorem op_inr {X Y Z : C} {f : X ⟶ Z} {g : Y ⟶ Z} (c : PullbackCone f g) : - c.op.inr = c.snd.op := by aesop_cat + c.op.inr = c.snd.op := by simp /-- If `c` is a pullback cone, then `c.op.unop` is isomorphic to `c`. -/ def opUnop {X Y Z : C} {f : X ⟶ Z} {g : Y ⟶ Z} (c : PullbackCone f g) : c.op.unop ≅ c := diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Basic.lean b/Mathlib/CategoryTheory/Limits/Preserves/Basic.lean index c77df5a351b5d..895a0d9129194 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Basic.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Basic.lean @@ -248,7 +248,7 @@ preservesLimit_of_preserves_limit_cone h hF lemma preservesLimit_of_iso_diagram {K₁ K₂ : J ⥤ C} (F : C ⥤ D) (h : K₁ ≅ K₂) [PreservesLimit K₁ F] : PreservesLimit K₂ F where preserves {c} t := ⟨by - apply IsLimit.postcomposeInvEquiv (isoWhiskerRight h F : _) _ _ + apply IsLimit.postcomposeInvEquiv (isoWhiskerRight h F :) _ _ have := (IsLimit.postcomposeInvEquiv h c).symm t apply IsLimit.ofIsoLimit (isLimitOfPreserves F this) exact Cones.ext (Iso.refl _)⟩ @@ -354,7 +354,7 @@ lemma preservesColimit_of_iso_diagram {K₁ K₂ : J ⥤ C} (F : C ⥤ D) (h : K [PreservesColimit K₁ F] : PreservesColimit K₂ F where preserves {c} t := ⟨by - apply IsColimit.precomposeHomEquiv (isoWhiskerRight h F : _) _ _ + apply IsColimit.precomposeHomEquiv (isoWhiskerRight h F :) _ _ have := (IsColimit.precomposeHomEquiv h c).symm t apply IsColimit.ofIsoColimit (isColimitOfPreserves F this) exact Cocones.ext (Iso.refl _)⟩ @@ -693,7 +693,7 @@ lemma reflectsLimit_of_iso_diagram {K₁ K₂ : J ⥤ C} (F : C ⥤ D) (h : K₁ ReflectsLimit K₂ F where reflects {c} t := ⟨by apply IsLimit.postcomposeInvEquiv h c (isLimitOfReflects F _) - apply ((IsLimit.postcomposeInvEquiv (isoWhiskerRight h F : _) _).symm t).ofIsoLimit _ + apply ((IsLimit.postcomposeInvEquiv (isoWhiskerRight h F :) _).symm t).ofIsoLimit _ exact Cones.ext (Iso.refl _)⟩ @[deprecated "use reflectsLimit_of_iso_diagram" (since := "2024-11-19")] @@ -868,7 +868,7 @@ lemma reflectsColimit_of_iso_diagram {K₁ K₂ : J ⥤ C} (F : C ⥤ D) (h : K ReflectsColimit K₂ F where reflects {c} t := ⟨by apply IsColimit.precomposeHomEquiv h c (isColimitOfReflects F _) - apply ((IsColimit.precomposeHomEquiv (isoWhiskerRight h F : _) _).symm t).ofIsoColimit _ + apply ((IsColimit.precomposeHomEquiv (isoWhiskerRight h F :) _).symm t).ofIsoColimit _ exact Cocones.ext (Iso.refl _)⟩ @[deprecated "use reflectsColimit_of_iso_diagram" (since := "2024-11-19")] diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Biproducts.lean b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Biproducts.lean index 550f8c8bea18d..95581e7fd2454 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Biproducts.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Biproducts.lean @@ -237,7 +237,7 @@ theorem biproductComparison'_comp_biproductComparison : @[simps] def splitEpiBiproductComparison : SplitEpi (biproductComparison F f) where section_ := biproductComparison' F f - id := by aesop + id := by simp instance : IsSplitEpi (biproductComparison F f) := IsSplitEpi.mk' (splitEpiBiproductComparison F f) @@ -246,7 +246,7 @@ instance : IsSplitEpi (biproductComparison F f) := @[simps] def splitMonoBiproductComparison' : SplitMono (biproductComparison' F f) where retraction := biproductComparison F f - id := by aesop + id := by simp instance : IsSplitMono (biproductComparison' F f) := IsSplitMono.mk' (splitMonoBiproductComparison' F f) @@ -321,7 +321,7 @@ theorem biprodComparison'_comp_biprodComparison : @[simps] def splitEpiBiprodComparison : SplitEpi (biprodComparison F X Y) where section_ := biprodComparison' F X Y - id := by aesop + id := by simp instance : IsSplitEpi (biprodComparison F X Y) := IsSplitEpi.mk' (splitEpiBiprodComparison F X Y) @@ -330,7 +330,7 @@ instance : IsSplitEpi (biprodComparison F X Y) := @[simps] def splitMonoBiprodComparison' : SplitMono (biprodComparison' F X Y) where retraction := biprodComparison F X Y - id := by aesop + id := by simp instance : IsSplitMono (biprodComparison' F X Y) := IsSplitMono.mk' (splitMonoBiprodComparison' F X Y) diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Pullbacks.lean b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Pullbacks.lean index 0eb7d4f49fdba..48b8cbe29f408 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Pullbacks.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Pullbacks.lean @@ -114,7 +114,7 @@ lemma preservesPullback_symmetry : PreservesLimit (cospan g f) G where infer_instance · exact (c.π.naturality WalkingCospan.Hom.inr).symm.trans - (c.π.naturality WalkingCospan.Hom.inl : _)⟩ + (c.π.naturality WalkingCospan.Hom.inl :)⟩ theorem hasPullback_of_preservesPullback [HasPullback f g] : HasPullback (G.map f) (G.map g) := ⟨⟨⟨_, isLimitPullbackConeMapOfIsLimit G _ (pullbackIsPullback _ _)⟩⟩⟩ diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Ulift.lean b/Mathlib/CategoryTheory/Limits/Preserves/Ulift.lean index d2998c6990db4..2b845755423e4 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Ulift.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Ulift.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2023 Dagur Asgeirsson. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Dagur Asgeirsson, Junyan Xu +Authors: Dagur Asgeirsson, Junyan Xu, Sophie Morel -/ import Mathlib.CategoryTheory.Limits.Creates import Mathlib.CategoryTheory.Limits.Types @@ -10,11 +10,12 @@ import Mathlib.Data.Set.Subsingleton /-! # `ULift` creates small (co)limits -This file shows that `uliftFunctor.{v, u}` creates (and hence preserves) limits and colimits indexed -by `J` with `[Category.{w, u} J]`. -It also shows that `uliftFunctor.{v, u}` *preserves* "all" limits, -potentially too big to exist in `Type u`). +This file shows that `uliftFunctor.{v, u}` preserves all limits and colimits, including those +potentially too big to exist in `Type u`. + +As this functor is fully faithful, we also deduce that it creates `u`-small limits and +colimits. -/ @@ -37,8 +38,7 @@ def sectionsEquiv {J : Type*} [Category J] (K : J ⥤ Type u) : /-- The functor `uliftFunctor : Type u ⥤ Type (max u v)` preserves limits of arbitrary size. -/ -noncomputable -instance : PreservesLimitsOfSize.{w', w} uliftFunctor.{v, u} where +noncomputable instance : PreservesLimitsOfSize.{w', w} uliftFunctor.{v, u} where preservesLimitsOfShape {J} := { preservesLimit := fun {K} => { preserves := fun {c} hc => by @@ -57,93 +57,16 @@ noncomputable instance : CreatesLimitsOfSize.{w, u} uliftFunctor.{v, u} where variable {J : Type*} [Category J] {K : J ⥤ Type u} {c : Cocone K} (hc : IsColimit c) variable {lc : Cocone (K ⋙ uliftFunctor.{v, u})} -/-- Given a subset of the cocone point of a cocone over the lifted functor, - produce a cocone over the original functor. -/ -def coconeOfSet (ls : Set lc.pt) : Cocone K where - pt := ULift Prop - ι := - { app := fun j x ↦ ⟨lc.ι.app j ⟨x⟩ ∈ ls⟩ - naturality := fun i j f ↦ by dsimp only; rw [← lc.w f]; rfl } - -/-- Given a subset of the cocone point of a cocone over the lifted functor, - produce a subset of the cocone point of a colimit cocone over the original functor. -/ -def descSet (ls : Set lc.pt) : Set c.pt := {x | (hc.desc (coconeOfSet ls) x).down} - -/-- Characterization the map `descSet hc`: the image of an element in a vertex of the original - diagram in the cocone point lies in `descSet hc ls` if and only if the image of the corresponding - element in the lifted diagram lie in `ls`. -/ -lemma descSet_spec (s : Set c.pt) (ls : Set lc.pt) : - descSet hc ls = s ↔ ∀ j x, lc.ι.app j ⟨x⟩ ∈ ls ↔ c.ι.app j x ∈ s := by - refine ⟨?_, fun he ↦ funext fun x ↦ ?_⟩ - · rintro rfl j x - exact (congr_arg ULift.down (congr_fun (hc.fac (coconeOfSet ls) j) x).symm).to_iff - · refine (congr_arg ULift.down (congr_fun (hc.uniq (coconeOfSet ls) (⟨· ∈ s⟩) fun j ↦ ?_) x)).symm - ext y; exact congr_arg ULift.up (propext (he j y).symm) - -lemma mem_descSet_singleton {x : lc.pt} {j : J} {y : K.obj j} : - c.ι.app j y ∈ descSet hc {x} ↔ lc.ι.app j ⟨y⟩ = x := - ((descSet_spec hc _ {x}).mp rfl j y).symm - -variable (lc) - -lemma descSet_univ : descSet hc (@Set.univ lc.pt) = Set.univ := by simp [descSet_spec] - -lemma iUnion_descSet_singleton : ⋃ x : lc.pt, descSet hc {x} = Set.univ := by - rw [← descSet_univ hc lc, eq_comm, descSet_spec] - intro j x - erw [true_iff, Set.mem_iUnion] - use lc.ι.app j ⟨x⟩ - rw [mem_descSet_singleton] - -lemma descSet_empty : descSet hc (∅ : Set lc.pt) = ∅ := by simp [descSet_spec] - -lemma descSet_inter_of_ne (x y : lc.pt) (hn : x ≠ y) : descSet hc {x} ∩ descSet hc {y} = ∅ := by - rw [eq_comm, ← descSet_empty hc lc, descSet_spec] - intro j z - erw [false_iff] - rintro ⟨hx, hy⟩ - rw [mem_descSet_singleton] at hx hy - exact hn (hx ▸ hy) - -lemma existsUnique_mem_descSet (x : c.pt) : ∃! y : lc.pt, x ∈ descSet hc {y} := - existsUnique_of_exists_of_unique - (Set.mem_iUnion.mp <| Set.eq_univ_iff_forall.mp (iUnion_descSet_singleton hc lc) x) - fun y₁ y₂ h₁ h₂ ↦ by_contra fun hn ↦ - Set.eq_empty_iff_forall_not_mem.1 (descSet_inter_of_ne hc lc y₁ y₂ hn) x ⟨h₁, h₂⟩ - -@[deprecated (since := "2024-12-17")] alias exists_unique_mem_descSet := existsUnique_mem_descSet - -/-- Given a colimit cocone in `Type u` and an arbitrary cocone over the diagram lifted to - `Type (max u v)`, produce a function from the cocone point of the colimit cocone to the - cocone point of the other cocone, that witnesses the colimit cocone also being a colimit - in the higher universe. -/ -noncomputable def descFun (x : c.pt) : lc.pt := (existsUnique_mem_descSet hc lc x).exists.choose - -lemma descFun_apply_spec {x : c.pt} {y : lc.pt} : descFun hc lc x = y ↔ x ∈ descSet hc {y} := - have hu := existsUnique_mem_descSet hc lc x - have hm := hu.exists.choose_spec - ⟨fun he ↦ he ▸ hm, hu.unique hm⟩ - -lemma descFun_spec (f : c.pt → lc.pt) : - f = descFun hc lc ↔ ∀ j, f ∘ c.ι.app j = lc.ι.app j ∘ ULift.up := by - refine ⟨?_, fun he ↦ funext fun x ↦ ((descFun_apply_spec hc lc).mpr ?_).symm⟩ - · rintro rfl j; ext - apply (descFun_apply_spec hc lc).mpr - rw [mem_descSet_singleton]; rfl - · rw [← (jointly_surjective_of_isColimit hc x).choose_spec.choose_spec, mem_descSet_singleton] - exact (congr_fun (he _) _).symm - /-- The functor `uliftFunctor : Type u ⥤ Type (max u v)` preserves colimits of arbitrary size. -/ noncomputable instance : PreservesColimitsOfSize.{w', w} uliftFunctor.{v, u} where preservesColimitsOfShape {J _} := { preservesColimit := fun {F} ↦ - { preserves := fun {c} hc ↦ ⟨{ - desc := fun lc x ↦ descFun hc lc x.down - fac := fun lc j ↦ by ext ⟨⟩; apply congr_fun ((descFun_spec hc lc _).mp rfl j) - uniq := fun lc f hf ↦ by ext ⟨⟩; apply congr_fun ((descFun_spec hc lc (f ∘ ULift.up)).mpr - fun j ↦ funext fun y ↦ congr_fun (hf j) ⟨y⟩) }⟩ } } + { preserves := fun {c} hc ↦ by + rw [isColimit_iff_bijective_desc, ← Function.Bijective.of_comp_iff _ + (quotQuotUliftEquiv F).bijective, Quot.desc_quotQuotUliftEquiv] + exact ULift.up_bijective.comp ((isColimit_iff_bijective_desc c).mp (Nonempty.intro hc)) } } /-- The functor `uliftFunctor : Type u ⥤ Type (max u v)` creates `u`-small colimits. diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean b/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean index cf91656226688..29d7bf37aa103 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean @@ -280,7 +280,7 @@ def whiskerToCone {f : J → C} (c : Bicone f) (g : K ≃ J) : (c.whisker g).toCone ≅ (Cones.postcompose (Discrete.functorComp f g).inv).obj (c.toCone.whisker (Discrete.functor (Discrete.mk ∘ g))) := - Cones.ext (Iso.refl _) (by aesop_cat) + Cones.ext (Iso.refl _) (by simp) /-- Taking the cocone of a whiskered bicone results in a cone isomorphic to one gained by whiskering the cocone and precomposing with a suitable isomorphism. -/ @@ -288,7 +288,7 @@ def whiskerToCocone {f : J → C} (c : Bicone f) (g : K ≃ J) : (c.whisker g).toCocone ≅ (Cocones.precompose (Discrete.functorComp f g).hom).obj (c.toCocone.whisker (Discrete.functor (Discrete.mk ∘ g))) := - Cocones.ext (Iso.refl _) (by aesop_cat) + Cocones.ext (Iso.refl _) (by simp) /-- Whiskering a bicone with an equivalence between types preserves being a bilimit bicone. -/ noncomputable def whiskerIsBilimitIff {f : J → C} (c : Bicone f) (g : K ≃ J) : @@ -579,7 +579,7 @@ theorem biproduct.map_eq_map' {f g : J → C} [HasBiproduct f] [HasBiproduct g] dsimp rw [biproduct.ι_π_assoc, biproduct.ι_π] split_ifs with h - · subst h; rw [eqToHom_refl, Category.id_comp]; erw [Category.comp_id] + · subst h; simp · simp @[reassoc (attr := simp)] @@ -625,7 +625,7 @@ instance biproduct.map_epi {f g : J → C} [HasBiproduct f] [HasBiproduct g] (p ι_colimMap_assoc, Discrete.functor_obj_eq_as, Discrete.natTrans_app, colimit.ι_desc_assoc, Cofan.mk_pt, Cofan.mk_ι_app, ι_π, ι_π_assoc] split - all_goals aesop + all_goals simp_all rw [this] infer_instance @@ -1757,7 +1757,7 @@ instance biprod.map_epi {W X Y Z : C} (f : W ⟶ Y) (g : X ⟶ Z) [Epi f] instance prod.map_epi {W X Y Z : C} (f : W ⟶ Y) (g : X ⟶ Z) [Epi f] [Epi g] [HasBinaryBiproduct W X] [HasBinaryBiproduct Y Z] : Epi (prod.map f g) := by rw [show prod.map f g = (biprod.isoProd _ _).inv ≫ biprod.map f g ≫ - (biprod.isoProd _ _).hom by aesop] + (biprod.isoProd _ _).hom by simp] infer_instance instance biprod.map_mono {W X Y Z : C} (f : W ⟶ Y) (g : X ⟶ Z) [Mono f] @@ -1769,7 +1769,7 @@ instance biprod.map_mono {W X Y Z : C} (f : W ⟶ Y) (g : X ⟶ Z) [Mono f] instance coprod.map_mono {W X Y Z : C} (f : W ⟶ Y) (g : X ⟶ Z) [Mono f] [Mono g] [HasBinaryBiproduct W X] [HasBinaryBiproduct Y Z] : Mono (coprod.map f g) := by rw [show coprod.map f g = (biprod.isoCoprod _ _).inv ≫ biprod.map f g ≫ - (biprod.isoCoprod _ _).hom by aesop] + (biprod.isoCoprod _ _).hom by simp] infer_instance section BiprodKernel diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Equalizers.lean b/Mathlib/CategoryTheory/Limits/Shapes/Equalizers.lean index 36a0a6e065433..fdf4259edc924 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Equalizers.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Equalizers.lean @@ -533,7 +533,7 @@ theorem Cofork.IsColimit.homIso_natural {X Y : C} {f g : X ⟶ Y} {t : Cofork f def Cone.ofFork {F : WalkingParallelPair ⥤ C} (t : Fork (F.map left) (F.map right)) : Cone F where pt := t.pt π := - { app := fun X => t.π.app X ≫ eqToHom (by aesop) + { app := fun X => t.π.app X ≫ eqToHom (by simp) naturality := by rintro _ _ (_|_|_) <;> {dsimp; simp [t.condition]}} /-- This is a helper construction that can be useful when verifying that a category has all @@ -548,23 +548,23 @@ def Cocone.ofCofork {F : WalkingParallelPair ⥤ C} (t : Cofork (F.map left) (F. Cocone F where pt := t.pt ι := - { app := fun X => eqToHom (by aesop) ≫ t.ι.app X + { app := fun X => eqToHom (by simp) ≫ t.ι.app X naturality := by rintro _ _ (_|_|_) <;> {dsimp; simp [t.condition]}} @[simp] theorem Cone.ofFork_π {F : WalkingParallelPair ⥤ C} (t : Fork (F.map left) (F.map right)) (j) : - (Cone.ofFork t).π.app j = t.π.app j ≫ eqToHom (by aesop) := rfl + (Cone.ofFork t).π.app j = t.π.app j ≫ eqToHom (by simp) := rfl @[simp] theorem Cocone.ofCofork_ι {F : WalkingParallelPair ⥤ C} (t : Cofork (F.map left) (F.map right)) - (j) : (Cocone.ofCofork t).ι.app j = eqToHom (by aesop) ≫ t.ι.app j := rfl + (j) : (Cocone.ofCofork t).ι.app j = eqToHom (by simp) ≫ t.ι.app j := rfl /-- Given `F : WalkingParallelPair ⥤ C`, which is really the same as `parallelPair (F.map left) (F.map right)` and a cone on `F`, we get a fork on `F.map left` and `F.map right`. -/ def Fork.ofCone {F : WalkingParallelPair ⥤ C} (t : Cone F) : Fork (F.map left) (F.map right) where pt := t.pt - π := { app := fun X => t.π.app X ≫ eqToHom (by aesop) + π := { app := fun X => t.π.app X ≫ eqToHom (by simp) naturality := by rintro _ _ (_|_|_) <;> {dsimp; simp}} /-- Given `F : WalkingParallelPair ⥤ C`, which is really the same as @@ -573,16 +573,16 @@ def Fork.ofCone {F : WalkingParallelPair ⥤ C} (t : Cone F) : Fork (F.map left) def Cofork.ofCocone {F : WalkingParallelPair ⥤ C} (t : Cocone F) : Cofork (F.map left) (F.map right) where pt := t.pt - ι := { app := fun X => eqToHom (by aesop) ≫ t.ι.app X + ι := { app := fun X => eqToHom (by simp) ≫ t.ι.app X naturality := by rintro _ _ (_|_|_) <;> {dsimp; simp}} @[simp] theorem Fork.ofCone_π {F : WalkingParallelPair ⥤ C} (t : Cone F) (j) : - (Fork.ofCone t).π.app j = t.π.app j ≫ eqToHom (by aesop) := rfl + (Fork.ofCone t).π.app j = t.π.app j ≫ eqToHom (by simp) := rfl @[simp] theorem Cofork.ofCocone_ι {F : WalkingParallelPair ⥤ C} (t : Cocone F) (j) : - (Cofork.ofCocone t).ι.app j = eqToHom (by aesop) ≫ t.ι.app j := rfl + (Cofork.ofCocone t).ι.app j = eqToHom (by simp) ≫ t.ι.app j := rfl @[simp] theorem Fork.ι_postcompose {f' g' : X ⟶ Y} {α : parallelPair f g ⟶ parallelPair f' g'} @@ -710,7 +710,7 @@ theorem equalizer.condition : equalizer.ι f g ≫ f = equalizer.ι f g ≫ g := /-- The equalizer built from `equalizer.ι f g` is limiting. -/ noncomputable def equalizerIsEqualizer : IsLimit (Fork.ofι (equalizer.ι f g) (equalizer.condition f g)) := - IsLimit.ofIsoLimit (limit.isLimit _) (Fork.ext (Iso.refl _) (by aesop)) + IsLimit.ofIsoLimit (limit.isLimit _) (Fork.ext (Iso.refl _) (by simp)) variable {f g} @@ -859,7 +859,7 @@ theorem coequalizer.condition : f ≫ coequalizer.π f g = g ≫ coequalizer.π /-- The cofork built from `coequalizer.π f g` is colimiting. -/ noncomputable def coequalizerIsCoequalizer : IsColimit (Cofork.ofπ (coequalizer.π f g) (coequalizer.condition f g)) := - IsColimit.ofIsoColimit (colimit.isColimit _) (Cofork.ext (Iso.refl _) (by aesop)) + IsColimit.ofIsoColimit (colimit.isColimit _) (Cofork.ext (Iso.refl _) (by simp)) variable {f g} diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Images.lean b/Mathlib/CategoryTheory/Limits/Shapes/Images.lean index 1214bfe40be59..f11792f5eaace 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Images.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Images.lean @@ -613,7 +613,7 @@ attribute [inherit_doc ImageMap] ImageMap.map ImageMap.map_ι attribute [-simp, nolint simpNF] ImageMap.mk.injEq instance inhabitedImageMap {f : Arrow C} [HasImage f.hom] : Inhabited (ImageMap (𝟙 f)) := - ⟨⟨𝟙 _, by aesop⟩⟩ + ⟨⟨𝟙 _, by simp⟩⟩ attribute [reassoc (attr := simp)] ImageMap.map_ι diff --git a/Mathlib/CategoryTheory/Limits/Shapes/IsTerminal.lean b/Mathlib/CategoryTheory/Limits/Shapes/IsTerminal.lean index de6ee324b9d3d..a33d65333e024 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/IsTerminal.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/IsTerminal.lean @@ -57,11 +57,11 @@ abbrev IsInitial (X : C) := /-- An object `Y` is terminal iff for every `X` there is a unique morphism `X ⟶ Y`. -/ def isTerminalEquivUnique (F : Discrete.{0} PEmpty.{1} ⥤ C) (Y : C) : - IsLimit (⟨Y, by aesop_cat, by aesop_cat⟩ : Cone F) ≃ ∀ X : C, Unique (X ⟶ Y) where + IsLimit (⟨Y, by aesop_cat, by simp⟩ : Cone F) ≃ ∀ X : C, Unique (X ⟶ Y) where toFun t X := - { default := t.lift ⟨X, ⟨by aesop_cat, by aesop_cat⟩⟩ + { default := t.lift ⟨X, ⟨by aesop_cat, by simp⟩⟩ uniq := fun f => - t.uniq ⟨X, ⟨by aesop_cat, by aesop_cat⟩⟩ f (by aesop_cat) } + t.uniq ⟨X, ⟨by aesop_cat, by simp⟩⟩ f (by simp) } invFun u := { lift := fun s => (u s.pt).default uniq := fun s _ _ => (u s.pt).2 _ } @@ -103,10 +103,10 @@ def IsTerminal.equivOfIso {X Y : C} (e : X ≅ Y) : /-- An object `X` is initial iff for every `Y` there is a unique morphism `X ⟶ Y`. -/ def isInitialEquivUnique (F : Discrete.{0} PEmpty.{1} ⥤ C) (X : C) : - IsColimit (⟨X, ⟨by aesop_cat, by aesop_cat⟩⟩ : Cocone F) ≃ ∀ Y : C, Unique (X ⟶ Y) where + IsColimit (⟨X, ⟨by aesop_cat, by simp⟩⟩ : Cocone F) ≃ ∀ Y : C, Unique (X ⟶ Y) where toFun t X := - { default := t.desc ⟨X, ⟨by aesop_cat, by aesop_cat⟩⟩ - uniq := fun f => t.uniq ⟨X, ⟨by aesop_cat, by aesop_cat⟩⟩ f (by aesop_cat) } + { default := t.desc ⟨X, ⟨by aesop_cat, by simp⟩⟩ + uniq := fun f => t.uniq ⟨X, ⟨by aesop_cat, by simp⟩⟩ f (by simp) } invFun u := { desc := fun s => (u s.pt).default uniq := fun s _ _ => (u s.pt).2 _ } @@ -152,7 +152,7 @@ def IsTerminal.from {X : C} (t : IsTerminal X) (Y : C) : Y ⟶ X := /-- Any two morphisms to a terminal object are equal. -/ theorem IsTerminal.hom_ext {X Y : C} (t : IsTerminal X) (f g : Y ⟶ X) : f = g := - IsLimit.hom_ext t (by aesop_cat) + IsLimit.hom_ext t (by simp) @[simp] theorem IsTerminal.comp_from {Z : C} (t : IsTerminal Z) {X Y : C} (f : X ⟶ Y) : @@ -169,7 +169,7 @@ def IsInitial.to {X : C} (t : IsInitial X) (Y : C) : X ⟶ Y := /-- Any two morphisms from an initial object are equal. -/ theorem IsInitial.hom_ext {X Y : C} (t : IsInitial X) (f g : X ⟶ Y) : f = g := - IsColimit.hom_ext t (by aesop_cat) + IsColimit.hom_ext t (by simp) @[simp] theorem IsInitial.to_comp {X : C} (t : IsInitial X) {Y Z : C} (f : Y ⟶ Z) : t.to Y ≫ f = t.to Z := @@ -217,12 +217,12 @@ variable (X : C) {F₁ : Discrete.{w} PEmpty ⥤ C} {F₂ : Discrete.{w'} PEmpty as long as the cone points are isomorphic. -/ def isLimitChangeEmptyCone {c₁ : Cone F₁} (hl : IsLimit c₁) (c₂ : Cone F₂) (hi : c₁.pt ≅ c₂.pt) : IsLimit c₂ where - lift c := hl.lift ⟨c.pt, by aesop_cat, by aesop_cat⟩ ≫ hi.hom + lift c := hl.lift ⟨c.pt, by aesop_cat, by simp⟩ ≫ hi.hom uniq c f _ := by dsimp rw [← hl.uniq _ (f ≫ hi.inv) _] · simp only [Category.assoc, Iso.inv_hom_id, Category.comp_id] - · aesop_cat + · simp /-- Replacing an empty cone in `IsLimit` by another with the same cone point is an equivalence. -/ @@ -239,12 +239,12 @@ def isLimitEmptyConeEquiv (c₁ : Cone F₁) (c₂ : Cone F₂) (h : c₁.pt ≅ as long as the cocone points are isomorphic. -/ def isColimitChangeEmptyCocone {c₁ : Cocone F₁} (hl : IsColimit c₁) (c₂ : Cocone F₂) (hi : c₁.pt ≅ c₂.pt) : IsColimit c₂ where - desc c := hi.inv ≫ hl.desc ⟨c.pt, by aesop_cat, by aesop_cat⟩ + desc c := hi.inv ≫ hl.desc ⟨c.pt, by aesop_cat, by simp⟩ uniq c f _ := by dsimp rw [← hl.uniq _ (hi.hom ≫ f) _] · simp only [Iso.inv_hom_id_assoc] - · aesop_cat + · simp /-- Replacing an empty cocone in `IsColimit` by another with the same cocone point is an equivalence. -/ diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Kernels.lean b/Mathlib/CategoryTheory/Limits/Shapes/Kernels.lean index 202f266d2d66f..6e11689c3fc8b 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Kernels.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Kernels.lean @@ -275,7 +275,7 @@ theorem kernel.condition : kernel.ι f ≫ f = 0 := /-- The kernel built from `kernel.ι f` is limiting. -/ def kernelIsKernel : IsLimit (Fork.ofι (kernel.ι f) ((kernel.condition f).trans comp_zero.symm)) := - IsLimit.ofIsoLimit (limit.isLimit _) (Fork.ext (Iso.refl _) (by aesop_cat)) + IsLimit.ofIsoLimit (limit.isLimit _) (Fork.ext (Iso.refl _) (by simp)) /-- Given any morphism `k : W ⟶ X` satisfying `k ≫ f = 0`, `k` factors through `kernel.ι f` via `kernel.lift : W ⟶ kernel f`. -/ @@ -425,8 +425,8 @@ instance hasKernel_iso_comp {X Y Z : C} (f : X ⟶ Y) (g : Y ⟶ Z) [IsIso f] [H HasKernel (f ≫ g) where exists_limit := ⟨{ cone := KernelFork.ofι (kernel.ι g ≫ inv f) (by simp) - isLimit := isLimitAux _ (fun s => kernel.lift _ (s.ι ≫ f) (by aesop_cat)) - (by aesop_cat) fun s m w => by + isLimit := isLimitAux _ (fun s => kernel.lift _ (s.ι ≫ f) (by simp)) + (by simp) fun s m w => by simp_rw [← w] symm apply equalizer.hom_ext @@ -751,7 +751,7 @@ lemma colimit_ι_zero_cokernel_desc {C : Type*} [Category C] [HasZeroMorphisms C] {X Y Z : C} (f : X ⟶ Y) (g : Y ⟶ Z) (h : f ≫ g = 0) [HasCokernel f] : colimit.ι (parallelPair f 0) WalkingParallelPair.zero ≫ cokernel.desc f g h = 0 := by rw [(colimit.w (parallelPair f 0) WalkingParallelPairHom.left).symm] - aesop_cat + simp @[simp] theorem cokernel.desc_zero {W : C} {h} : cokernel.desc f (0 : Y ⟶ W) h = 0 := by @@ -884,7 +884,7 @@ instance hasCokernel_comp_iso {X Y Z : C} (f : X ⟶ Y) (g : Y ⟶ Z) [HasCokern isColimitAux _ (fun s => cokernel.desc _ (g ≫ s.π) (by rw [← Category.assoc, CokernelCofork.condition])) - (by aesop_cat) fun s m w => by + (by simp) fun s m w => by simp_rw [← w] symm apply coequalizer.hom_ext diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Multiequalizer.lean b/Mathlib/CategoryTheory/Limits/Shapes/Multiequalizer.lean index 83399224cf13d..b61c4579338d9 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Multiequalizer.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Multiequalizer.lean @@ -586,8 +586,8 @@ noncomputable def ofSigmaCofork (c : Cofork I.fstSigmaMap I.sndSigmaMap) : Multi ι := { app := fun x => match x with - | WalkingMultispan.left a => (Sigma.ι I.left a : _) ≫ I.fstSigmaMap ≫ c.π - | WalkingMultispan.right b => (Sigma.ι I.right b : _) ≫ c.π + | WalkingMultispan.left a => (Sigma.ι I.left a :) ≫ I.fstSigmaMap ≫ c.π + | WalkingMultispan.right b => (Sigma.ι I.right b :) ≫ c.π naturality := by rintro (_ | _) (_ | _) (_ | _ | _) · simp @@ -600,17 +600,17 @@ noncomputable def ofSigmaCofork (c : Cofork I.fstSigmaMap I.sndSigmaMap) : Multi @[simp] theorem ofSigmaCofork_ι_app_left (c : Cofork I.fstSigmaMap I.sndSigmaMap) (a) : (ofSigmaCofork I c).ι.app (WalkingMultispan.left a) = - (Sigma.ι I.left a : _) ≫ I.fstSigmaMap ≫ c.π := + (Sigma.ι I.left a :) ≫ I.fstSigmaMap ≫ c.π := rfl -- @[simp] -- Porting note: LHS simplifies to obtain the normal form below theorem ofSigmaCofork_ι_app_right (c : Cofork I.fstSigmaMap I.sndSigmaMap) (b) : - (ofSigmaCofork I c).ι.app (WalkingMultispan.right b) = (Sigma.ι I.right b : _) ≫ c.π := + (ofSigmaCofork I c).ι.app (WalkingMultispan.right b) = (Sigma.ι I.right b :) ≫ c.π := rfl @[simp] theorem ofSigmaCofork_ι_app_right' (c : Cofork I.fstSigmaMap I.sndSigmaMap) (b) : - π (ofSigmaCofork I c) b = (Sigma.ι I.right b : _) ≫ c.π := + π (ofSigmaCofork I c) b = (Sigma.ι I.right b :) ≫ c.π := rfl end Multicofork diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Products.lean b/Mathlib/CategoryTheory/Limits/Shapes/Products.lean index 02cc099007693..9cd40dec3a39b 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Products.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Products.lean @@ -565,7 +565,7 @@ instance {ι : Type*} (f : ι → Type*) (g : (i : ι) → (f i) → C) exists_limit := Nonempty.intro { cone := Fan.mk (∏ᶜ fun i => ∏ᶜ g i) (fun X => Pi.π (fun i => ∏ᶜ g i) X.1 ≫ Pi.π (g X.1) X.2) isLimit := mkFanLimit _ (fun s => Pi.lift fun b => Pi.lift fun c => s.proj ⟨b, c⟩) - (by aesop_cat) + (by simp) (by intro s m w; simp only [Fan.mk_pt]; symm; ext i x; simp_all [Sigma.forall]) } /-- An iterated product is a product over a sigma type. -/ @@ -586,7 +586,7 @@ instance {ι : Type*} (f : ι → Type*) (g : (i : ι) → (f i) → C) (fun X => Sigma.ι (g X.1) X.2 ≫ Sigma.ι (fun i => ∐ g i) X.1) isColimit := mkCofanColimit _ (fun s => Sigma.desc fun b => Sigma.desc fun c => s.inj ⟨b, c⟩) - (by aesop_cat) + (by simp) (by intro s m w; simp only [Cofan.mk_pt]; symm; ext i x; simp_all [Sigma.forall]) } /-- An iterated coproduct is a coproduct over a sigma type. -/ diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Pullback/CommSq.lean b/Mathlib/CategoryTheory/Limits/Shapes/Pullback/CommSq.lean index 14f78f89b5b08..bc2c51f174f0e 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Pullback/CommSq.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Pullback/CommSq.lean @@ -89,27 +89,27 @@ theorem cocone_inr (s : CommSq f g h i) : s.cocone.inr = i := a commutative square identifies to the cocone of the flipped commutative square in the opposite category -/ def coneOp (p : CommSq f g h i) : p.cone.op ≅ p.flip.op.cocone := - PushoutCocone.ext (Iso.refl _) (by aesop_cat) (by aesop_cat) + PushoutCocone.ext (Iso.refl _) (by simp) (by simp) /-- The pullback cone in the opposite category associated to the cocone of a commutative square identifies to the cone of the flipped commutative square in the opposite category -/ def coconeOp (p : CommSq f g h i) : p.cocone.op ≅ p.flip.op.cone := - PullbackCone.ext (Iso.refl _) (by aesop_cat) (by aesop_cat) + PullbackCone.ext (Iso.refl _) (by simp) (by simp) /-- The pushout cocone obtained from the pullback cone associated to a commutative square in the opposite category identifies to the cocone associated to the flipped square. -/ def coneUnop {W X Y Z : Cᵒᵖ} {f : W ⟶ X} {g : W ⟶ Y} {h : X ⟶ Z} {i : Y ⟶ Z} (p : CommSq f g h i) : p.cone.unop ≅ p.flip.unop.cocone := - PushoutCocone.ext (Iso.refl _) (by aesop_cat) (by aesop_cat) + PushoutCocone.ext (Iso.refl _) (by simp) (by simp) /-- The pullback cone obtained from the pushout cone associated to a commutative square in the opposite category identifies to the cone associated to the flipped square. -/ def coconeUnop {W X Y Z : Cᵒᵖ} {f : W ⟶ X} {g : W ⟶ Y} {h : X ⟶ Z} {i : Y ⟶ Z} (p : CommSq f g h i) : p.cocone.unop ≅ p.flip.unop.cone := - PullbackCone.ext (Iso.refl _) (by aesop_cat) (by aesop_cat) + PullbackCone.ext (Iso.refl _) (by simp) (by simp) end CommSq @@ -221,7 +221,7 @@ lemma hom_ext (hP : IsPullback fst snd f g) {W : C} {k l : W ⟶ P} theorem of_isLimit {c : PullbackCone f g} (h : Limits.IsLimit c) : IsPullback c.fst c.snd f g := { w := c.condition isLimit' := ⟨IsLimit.ofIsoLimit h (Limits.PullbackCone.ext (Iso.refl _) - (by aesop_cat) (by aesop_cat))⟩ } + (by simp) (by simp))⟩ } /-- A variant of `of_isLimit` that is more useful with `apply`. -/ theorem of_isLimit' (w : CommSq fst snd f g) (h : Limits.IsLimit w.cone) : @@ -345,7 +345,7 @@ theorem of_horiz_isIso [IsIso fst] [IsIso g] (sq : CommSq fst snd f g) : IsPullb of_isLimit' sq (by refine - PullbackCone.IsLimit.mk _ (fun s => s.fst ≫ inv fst) (by aesop_cat) + PullbackCone.IsLimit.mk _ (fun s => s.fst ≫ inv fst) (by simp) (fun s => ?_) (by aesop_cat) simp only [← cancel_mono g, Category.assoc, ← sq.w, IsIso.inv_hom_id_assoc, s.condition]) @@ -431,7 +431,7 @@ theorem of_isColimit {c : PushoutCocone f g} (h : Limits.IsColimit c) : IsPushou { w := c.condition isColimit' := ⟨IsColimit.ofIsoColimit h (Limits.PushoutCocone.ext (Iso.refl _) - (by aesop_cat) (by aesop_cat))⟩ } + (by simp) (by simp))⟩ } /-- A variant of `of_isColimit` that is more useful with `apply`. -/ theorem of_isColimit' (w : CommSq f g inl inr) (h : Limits.IsColimit w.cocone) : @@ -1161,7 +1161,7 @@ theorem of_horiz_isIso [IsIso f] [IsIso inr] (sq : CommSq f g inl inr) : IsPusho (by refine PushoutCocone.IsColimit.mk _ (fun s => inv inr ≫ s.inr) (fun s => ?_) - (by aesop_cat) (by aesop_cat) + (by simp) (by simp) simp only [← cancel_epi f, s.condition, sq.w_assoc, IsIso.hom_inv_id_assoc]) theorem of_vert_isIso [IsIso g] [IsIso inl] (sq : CommSq f g inl inr) : IsPushout f g inl inr := diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Pullback/Mono.lean b/Mathlib/CategoryTheory/Limits/Shapes/Pullback/Mono.lean index 5b6bfbf8ab8b4..315be3a49b8d5 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Pullback/Mono.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Pullback/Mono.lean @@ -112,7 +112,7 @@ def isLimitOfCompMono (f : X ⟶ W) (g : Y ⟶ W) (i : W ⟶ Z) [Mono i] (s : Pu ⟨l, h₁, h₂⟩ refine ⟨l, h₁, h₂, ?_⟩ intro m hm₁ hm₂ - exact (PullbackCone.IsLimit.hom_ext H (hm₁.trans h₁.symm) (hm₂.trans h₂.symm) : _) + exact (PullbackCone.IsLimit.hom_ext H (hm₁.trans h₁.symm) (hm₂.trans h₂.symm) :) end PullbackCone @@ -299,7 +299,7 @@ def isColimitOfEpiComp (f : X ⟶ Y) (g : X ⟶ Z) (h : W ⟶ X) [Epi h] (s : Pu ⟨l, h₁, h₂⟩ refine ⟨l, h₁, h₂, ?_⟩ intro m hm₁ hm₂ - exact (PushoutCocone.IsColimit.hom_ext H (hm₁.trans h₁.symm) (hm₂.trans h₂.symm) : _) + exact (PushoutCocone.IsColimit.hom_ext H (hm₁.trans h₁.symm) (hm₂.trans h₂.symm) :) end PushoutCocone diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Pullback/PullbackCone.lean b/Mathlib/CategoryTheory/Limits/Shapes/Pullback/PullbackCone.lean index 0c26846165918..04011c1c45cb3 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Pullback/PullbackCone.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Pullback/PullbackCone.lean @@ -286,7 +286,7 @@ composing with `diagramIsoCospan`. -/ def PullbackCone.isoMk {F : WalkingCospan ⥤ C} (t : Cone F) : (Cones.postcompose (diagramIsoCospan.{v} _).hom).obj t ≅ PullbackCone.mk (t.π.app WalkingCospan.left) (t.π.app WalkingCospan.right) - ((t.π.naturality inl).symm.trans (t.π.naturality inr : _)) := + ((t.π.naturality inl).symm.trans (t.π.naturality inr :)) := Cones.ext (Iso.refl _) <| by rintro (_ | (_ | _)) <;> · dsimp diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Pullback/Square.lean b/Mathlib/CategoryTheory/Limits/Shapes/Pullback/Square.lean index 01bb011074096..ddd0fc6d43918 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Pullback/Square.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Pullback/Square.lean @@ -79,7 +79,7 @@ lemma IsPullback.of_iso {sq₁ sq₂ : Square C} (h : sq₁.IsPullback) refine CategoryTheory.IsPullback.of_iso h (evaluation₁.mapIso e) (evaluation₂.mapIso e) (evaluation₃.mapIso e) (evaluation₄.mapIso e) ?_ ?_ ?_ ?_ - all_goals aesop_cat + all_goals simp lemma IsPullback.iff_of_iso {sq₁ sq₂ : Square C} (e : sq₁ ≅ sq₂) : sq₁.IsPullback ↔ sq₂.IsPullback := @@ -90,7 +90,7 @@ lemma IsPushout.of_iso {sq₁ sq₂ : Square C} (h : sq₁.IsPushout) refine CategoryTheory.IsPushout.of_iso h (evaluation₁.mapIso e) (evaluation₂.mapIso e) (evaluation₃.mapIso e) (evaluation₄.mapIso e) ?_ ?_ ?_ ?_ - all_goals aesop_cat + all_goals simp lemma IsPushout.iff_of_iso {sq₁ sq₂ : Square C} (e : sq₁ ≅ sq₂) : sq₁.IsPushout ↔ sq₂.IsPushout := diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Reflexive.lean b/Mathlib/CategoryTheory/Limits/Shapes/Reflexive.lean index abcc9d116cda0..d1c5d29b44ebf 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Reflexive.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Reflexive.lean @@ -543,7 +543,7 @@ lemma reflexiveCoforkEquivCofork_functor_obj_π (G : ReflexiveCofork F) : ((reflexiveCoforkEquivCofork F).functor.obj G).π = G.π := by dsimp [reflexiveCoforkEquivCofork] rw [ReflexiveCofork.π, Cofork.π] - aesop_cat + simp @[simp] lemma reflexiveCoforkEquivCofork_inverse_obj_π diff --git a/Mathlib/CategoryTheory/Limits/Shapes/StrongEpi.lean b/Mathlib/CategoryTheory/Limits/Shapes/StrongEpi.lean index 2abdbdfda2ec4..e5f5d8bc1363b 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/StrongEpi.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/StrongEpi.lean @@ -164,11 +164,11 @@ end /-- A strong epimorphism that is a monomorphism is an isomorphism. -/ theorem isIso_of_mono_of_strongEpi (f : P ⟶ Q) [Mono f] [StrongEpi f] : IsIso f := - ⟨⟨(CommSq.mk (show 𝟙 P ≫ f = f ≫ 𝟙 Q by simp)).lift, by aesop_cat⟩⟩ + ⟨⟨(CommSq.mk (show 𝟙 P ≫ f = f ≫ 𝟙 Q by simp)).lift, by simp⟩⟩ /-- A strong monomorphism that is an epimorphism is an isomorphism. -/ theorem isIso_of_epi_of_strongMono (f : P ⟶ Q) [Epi f] [StrongMono f] : IsIso f := - ⟨⟨(CommSq.mk (show 𝟙 P ≫ f = f ≫ 𝟙 Q by simp)).lift, by aesop_cat⟩⟩ + ⟨⟨(CommSq.mk (show 𝟙 P ≫ f = f ≫ 𝟙 Q by simp)).lift, by simp⟩⟩ section diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Terminal.lean b/Mathlib/CategoryTheory/Limits/Shapes/Terminal.lean index d24c05e73cd2c..e3f907ea79b4a 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Terminal.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Terminal.lean @@ -43,7 +43,7 @@ section Univ variable (X : C) {F₁ : Discrete.{w} PEmpty ⥤ C} {F₂ : Discrete.{w'} PEmpty ⥤ C} theorem hasTerminalChangeDiagram (h : HasLimit F₁) : HasLimit F₂ := - ⟨⟨⟨⟨limit F₁, by aesop_cat, by aesop_cat⟩, + ⟨⟨⟨⟨limit F₁, by aesop_cat, by simp⟩, isLimitChangeEmptyCone C (limit.isLimit F₁) _ (eqToIso rfl)⟩⟩⟩ theorem hasTerminalChangeUniverse [h : HasLimitsOfShape (Discrete.{w} PEmpty) C] : @@ -51,7 +51,7 @@ theorem hasTerminalChangeUniverse [h : HasLimitsOfShape (Discrete.{w} PEmpty) C] has_limit _ := hasTerminalChangeDiagram C (h.1 (Functor.empty C)) theorem hasInitialChangeDiagram (h : HasColimit F₁) : HasColimit F₂ := - ⟨⟨⟨⟨colimit F₁, by aesop_cat, by aesop_cat⟩, + ⟨⟨⟨⟨colimit F₁, by aesop_cat, by simp⟩, isColimitChangeEmptyCocone C (colimit.isColimit F₁) _ (eqToIso rfl)⟩⟩⟩ theorem hasInitialChangeUniverse [h : HasColimitsOfShape (Discrete.{w} PEmpty) C] : @@ -92,7 +92,7 @@ theorem hasTerminal_of_unique (X : C) [∀ Y, Nonempty (Y ⟶ X)] [∀ Y, Subsin ⟨Classical.inhabited_of_nonempty', (Subsingleton.elim · _)⟩⟩ theorem IsTerminal.hasTerminal {X : C} (h : IsTerminal X) : HasTerminal C := - { has_limit := fun F => HasLimit.mk ⟨⟨X, by aesop_cat, by aesop_cat⟩, + { has_limit := fun F => HasLimit.mk ⟨⟨X, by aesop_cat, by simp⟩, isLimitChangeEmptyCone _ h _ (Iso.refl _)⟩ } /-- We can more explicitly show that a category has an initial object by specifying the object, @@ -104,7 +104,7 @@ theorem hasInitial_of_unique (X : C) [∀ Y, Nonempty (X ⟶ Y)] [∀ Y, Subsing theorem IsInitial.hasInitial {X : C} (h : IsInitial X) : HasInitial C where has_colimit F := - HasColimit.mk ⟨⟨X, by aesop_cat, by aesop_cat⟩, isColimitChangeEmptyCocone _ h _ (Iso.refl _)⟩ + HasColimit.mk ⟨⟨X, by aesop_cat, by simp⟩, isColimitChangeEmptyCocone _ h _ (Iso.refl _)⟩ /-- The map from an object to the terminal object. -/ abbrev terminal.from [HasTerminal C] (P : C) : P ⟶ ⊤_ C := diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Types.lean b/Mathlib/CategoryTheory/Limits/Shapes/Types.lean index f6dfbe28fd6cb..a7e173c76621e 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Types.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Types.lean @@ -346,7 +346,7 @@ def productLimitCone {J : Type v} (F : J → TypeMax.{v, u}) : π := Discrete.natTrans (fun ⟨j⟩ f => f j) } isLimit := { lift := fun s x j => s.π.app ⟨j⟩ x - uniq := fun _ _ w => funext fun x => funext fun j => (congr_fun (w ⟨j⟩) x : _) } + uniq := fun _ _ w => funext fun x => funext fun j => (congr_fun (w ⟨j⟩) x :) } /-- The categorical product in `TypeMax.{v, u}` is the type theoretic product `Π j, F j`. -/ noncomputable def productIso {J : Type v} (F : J → TypeMax.{v, u}) : ∏ᶜ F ≅ ∀ j, F j := @@ -385,7 +385,7 @@ noncomputable def productLimitCone : have : Small.{u} (∀ j, F j) := inferInstance { lift := fun s x => (equivShrink _) (fun j => s.π.app ⟨j⟩ x) uniq := fun s m w => funext fun x => Shrink.ext <| funext fun j => by - simpa using (congr_fun (w ⟨j⟩) x : _) } + simpa using (congr_fun (w ⟨j⟩) x :) } /-- The categorical product in `Type u` indexed in `Type v` is the type theoretic product `Π j, F j`, after shrinking back to `Type u`. -/ @@ -924,7 +924,7 @@ lemma pushoutCocone_inl_eq_inr_iff_of_isColimit {c : PushoutCocone f g} (hc : Is c.inl x₁ = c.inr x₂ ↔ ∃ (s : S), f s = x₁ ∧ g s = x₂ := by rw [pushoutCocone_inl_eq_inr_iff_of_iso (Cocones.ext (IsColimit.coconePointUniqueUpToIso hc (Pushout.isColimitCocone f g)) - (by aesop_cat))] + (by simp))] have := (mono_iff_injective f).2 h₁ apply Pushout.inl_eq_inr_iff diff --git a/Mathlib/CategoryTheory/Limits/Shapes/WidePullbacks.lean b/Mathlib/CategoryTheory/Limits/Shapes/WidePullbacks.lean index 4670206841758..06d17c50c7d91 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/WidePullbacks.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/WidePullbacks.lean @@ -330,9 +330,9 @@ theorem eq_lift_of_comp_eq (g : X ⟶ widePullback _ _ arrows) : · apply h1 theorem hom_eq_lift (g : X ⟶ widePullback _ _ arrows) : - g = lift (g ≫ base arrows) (fun j => g ≫ π arrows j) (by aesop_cat) := by + g = lift (g ≫ base arrows) (fun j => g ≫ π arrows j) (by simp) := by apply eq_lift_of_comp_eq - · aesop_cat + · simp · rfl -- Porting note: quite a few missing refl's in aesop_cat now @[ext 1100] @@ -400,7 +400,7 @@ theorem hom_eq_desc (g : widePushout _ _ arrows ⟶ X) : rw [← Category.assoc] simp := by apply eq_desc_of_comp_eq - · aesop_cat + · simp · rfl -- Porting note: another missing rfl @[ext 1100] diff --git a/Mathlib/CategoryTheory/Limits/Types.lean b/Mathlib/CategoryTheory/Limits/Types.lean index 0eff7c3add87e..fd38583178553 100644 --- a/Mathlib/CategoryTheory/Limits/Types.lean +++ b/Mathlib/CategoryTheory/Limits/Types.lean @@ -23,7 +23,7 @@ and the type `lim Hom(F·, X)`. open CategoryTheory CategoryTheory.Limits -universe v u w +universe u' v u w namespace CategoryTheory.Limits @@ -357,6 +357,56 @@ lemma Quot.ι_desc (j : J) (x : F.obj j) : Quot.desc c (Quot.ι F j x) = c.ι.ap lemma Quot.map_ι {j j' : J} {f : j ⟶ j'} (x : F.obj j) : Quot.ι F j' (F.map f x) = Quot.ι F j x := (Quot.sound ⟨f, rfl⟩).symm +/-- +The obvious map from `Quot F` to `Quot (F ⋙ uliftFunctor.{u'})`. +-/ +def quotToQuotUlift (F : J ⥤ Type u) : Quot F → Quot (F ⋙ uliftFunctor.{u'}) := by + refine Quot.lift (fun ⟨j, x⟩ ↦ Quot.ι _ j (ULift.up x)) ?_ + intro ⟨j, x⟩ ⟨j', y⟩ ⟨(f : j ⟶ j'), (eq : y = F.map f x)⟩ + dsimp + have eq : ULift.up y = (F ⋙ uliftFunctor.{u'}).map f (ULift.up x) := by + rw [eq] + dsimp + rw [eq, Quot.map_ι] + +@[simp] +lemma quotToQuotUlift_ι (F : J ⥤ Type u) (j : J) (x : F.obj j) : + quotToQuotUlift F (Quot.ι F j x) = Quot.ι _ j (ULift.up x) := by + dsimp [quotToQuotUlift, Quot.ι] + +/-- +The obvious map from `Quot (F ⋙ uliftFunctor.{u'})` to `Quot F`. +-/ +def quotUliftToQuot (F : J ⥤ Type u) : Quot (F ⋙ uliftFunctor.{u'}) → Quot F := + Quot.lift (fun ⟨j, x⟩ ↦ Quot.ι _ j x.down) + (fun ⟨_, x⟩ ⟨_, y⟩ ⟨f, (eq : y = ULift.up (F.map f x.down))⟩ ↦ by simp [eq, Quot.map_ι]) + +@[simp] +lemma quotUliftToQuot_ι (F : J ⥤ Type u) (j : J) (x : (F ⋙ uliftFunctor.{u'}).obj j) : + quotUliftToQuot F (Quot.ι _ j x) = Quot.ι F j x.down := by + dsimp [quotUliftToQuot, Quot.ι] + +/-- +The equivalence between `Quot F` and `Quot (F ⋙ uliftFunctor.{u'})`. +-/ +@[simp] +def quotQuotUliftEquiv (F : J ⥤ Type u) : Quot F ≃ Quot (F ⋙ uliftFunctor.{u'}) where + toFun := quotToQuotUlift F + invFun := quotUliftToQuot F + left_inv x := by + obtain ⟨j, y, rfl⟩ := Quot.jointly_surjective x + rw [quotToQuotUlift_ι, quotUliftToQuot_ι] + right_inv x := by + obtain ⟨j, y, rfl⟩ := Quot.jointly_surjective x + rw [quotUliftToQuot_ι, quotToQuotUlift_ι] + rfl + +lemma Quot.desc_quotQuotUliftEquiv {F : J ⥤ Type u} (c : Cocone F) : + Quot.desc (uliftFunctor.{u'}.mapCocone c) ∘ quotQuotUliftEquiv F = ULift.up ∘ Quot.desc c := by + ext x + obtain ⟨_, _, rfl⟩ := Quot.jointly_surjective x + dsimp + /-- (implementation detail) A function `Quot F → α` induces a cocone on `F` as long as the universes work out. -/ @[simps] diff --git a/Mathlib/CategoryTheory/Limits/VanKampen.lean b/Mathlib/CategoryTheory/Limits/VanKampen.lean index 705c13f505a00..f66b56e21c46b 100644 --- a/Mathlib/CategoryTheory/Limits/VanKampen.lean +++ b/Mathlib/CategoryTheory/Limits/VanKampen.lean @@ -105,10 +105,10 @@ theorem IsVanKampenColimit.isUniversal {F : J ⥤ C} {c : Cocone F} (H : IsVanKa /-- A universal colimit is a colimit. -/ noncomputable def IsUniversalColimit.isColimit {F : J ⥤ C} {c : Cocone F} (h : IsUniversalColimit c) : IsColimit c := by - refine ((h c (𝟙 F) (𝟙 c.pt : _) (by rw [Functor.map_id, Category.comp_id, Category.id_comp]) + refine ((h c (𝟙 F) (𝟙 c.pt :) (by rw [Functor.map_id, Category.comp_id, Category.id_comp]) (NatTrans.equifibered_of_isIso _)) fun j => ?_).some haveI : IsIso (𝟙 c.pt) := inferInstance - exact IsPullback.of_vert_isIso ⟨by erw [NatTrans.id_app, Category.comp_id, Category.id_comp]⟩ + exact IsPullback.of_vert_isIso ⟨by simp⟩ /-- A van Kampen colimit is a colimit. -/ noncomputable def IsVanKampenColimit.isColimit {F : J ⥤ C} {c : Cocone F} @@ -345,10 +345,7 @@ theorem IsUniversalColimit.map_reflective let cf : (Cocones.precompose β.hom).obj c' ⟶ Gl.mapCocone c'' := by refine { hom := pullback.lift ?_ f ?_ ≫ (PreservesPullback.iso _ _ _).inv, w := ?_ } · exact inv <| adj.counit.app c'.pt - · rw [IsIso.inv_comp_eq, ← adj.counit_naturality_assoc f, ← cancel_mono (adj.counit.app <| - Gl.obj c.pt), Category.assoc, Category.assoc, adj.left_triangle_components] - erw [Category.comp_id] - rfl + · simp [← cancel_mono (adj.counit.app <| Gl.obj c.pt)] · intro j rw [← Category.assoc, Iso.comp_inv_eq] ext @@ -426,7 +423,7 @@ theorem IsVanKampenColimit.map_reflective [HasColimitsOfShape J C] Gl.map (colimit.desc _ ⟨_, whiskerRight α' Gr ≫ c.2⟩) := by symm convert @IsColimit.coconePointUniqueUpToIso_hom_desc _ _ _ _ ((F' ⋙ Gr) ⋙ Gl) - (Gl.mapCocone ⟨_, (whiskerRight α' Gr ≫ c.2 : _)⟩) _ _ hl hr using 2 + (Gl.mapCocone ⟨_, (whiskerRight α' Gr ≫ c.2 :)⟩) _ _ hl hr using 2 · apply hr.hom_ext intro j rw [hr.fac, Functor.mapCocone_ι_app, ← Gl.map_comp, colimit.cocone_ι, colimit.ι_desc] diff --git a/Mathlib/CategoryTheory/Localization/Bifunctor.lean b/Mathlib/CategoryTheory/Localization/Bifunctor.lean index e4556b6e608d8..7755ae9cca22e 100644 --- a/Mathlib/CategoryTheory/Localization/Bifunctor.lean +++ b/Mathlib/CategoryTheory/Localization/Bifunctor.lean @@ -168,8 +168,8 @@ and `Lifting₂ L₁ L₂ W₁ W₂ F₂ F₂'` hold. -/ noncomputable def lift₂NatIso (e : F₁ ≅ F₂) : F₁' ≅ F₂' where hom := lift₂NatTrans L₁ L₂ W₁ W₂ F₁ F₂ F₁' F₂' e.hom inv := lift₂NatTrans L₁ L₂ W₁ W₂ F₂ F₁ F₂' F₁' e.inv - hom_inv_id := natTrans₂_ext L₁ L₂ W₁ W₂ (by aesop_cat) - inv_hom_id := natTrans₂_ext L₁ L₂ W₁ W₂ (by aesop_cat) + hom_inv_id := natTrans₂_ext L₁ L₂ W₁ W₂ (by simp) + inv_hom_id := natTrans₂_ext L₁ L₂ W₁ W₂ (by simp) end diff --git a/Mathlib/CategoryTheory/Localization/Construction.lean b/Mathlib/CategoryTheory/Localization/Construction.lean index b82770a4d5037..5bcd82a33e3e7 100644 --- a/Mathlib/CategoryTheory/Localization/Construction.lean +++ b/Mathlib/CategoryTheory/Localization/Construction.lean @@ -154,7 +154,7 @@ def lift : W.Localization ⥤ D := -- Porting note: rest of proof was `rcases r with ⟨⟩; tidy` rcases r with (_|_|⟨f,hf⟩|⟨f,hf⟩) · aesop_cat - · aesop_cat + · simp all_goals dsimp haveI := hG f hf diff --git a/Mathlib/CategoryTheory/Localization/DerivabilityStructure/Constructor.lean b/Mathlib/CategoryTheory/Localization/DerivabilityStructure/Constructor.lean index 2adefbca8f183..b7da168c46084 100644 --- a/Mathlib/CategoryTheory/Localization/DerivabilityStructure/Constructor.lean +++ b/Mathlib/CategoryTheory/Localization/DerivabilityStructure/Constructor.lean @@ -90,7 +90,7 @@ lemma isConnected : (isPreconnected_zigzag (RightResolution.mk (𝟙 _) (W₂.id_mem _)) (RightResolution.mk ρ.w.right ρ.hw.2)) refine Zigzag.trans ?_ (Zigzag.trans this ?_) - · exact Zigzag.of_hom (eqToHom (by aesop)) + · exact Zigzag.of_hom (eqToHom (by simp)) · apply Zigzag.of_inv refine CostructuredArrow.homMk (StructuredArrow.homMk ρ.X₁.hom (by simp)) ?_ ext diff --git a/Mathlib/CategoryTheory/Monad/Adjunction.lean b/Mathlib/CategoryTheory/Monad/Adjunction.lean index 7c5e8d990bb5a..447df18a698b9 100644 --- a/Mathlib/CategoryTheory/Monad/Adjunction.lean +++ b/Mathlib/CategoryTheory/Monad/Adjunction.lean @@ -172,7 +172,7 @@ theorem Monad.left_comparison (h : L ⊣ R) : L ⋙ Monad.comparison h = h.toMon rfl instance [R.Faithful] (h : L ⊣ R) : (Monad.comparison h).Faithful where - map_injective {_ _} _ _ w := R.map_injective (congr_arg Monad.Algebra.Hom.f w : _) + map_injective {_ _} _ _ w := R.map_injective (congr_arg Monad.Algebra.Hom.f w :) instance (T : Monad C) : (Monad.comparison T.adj).Full where map_surjective {_ _} f := ⟨⟨f.f, by simpa using f.h⟩, rfl⟩ @@ -218,7 +218,7 @@ theorem Comonad.left_comparison (h : L ⊣ R) : R ⋙ Comonad.comparison h = h.t instance Comonad.comparison_faithful_of_faithful [L.Faithful] (h : L ⊣ R) : (Comonad.comparison h).Faithful where - map_injective {_ _} _ _ w := L.map_injective (congr_arg Comonad.Coalgebra.Hom.f w : _) + map_injective {_ _} _ _ w := L.map_injective (congr_arg Comonad.Coalgebra.Hom.f w :) instance (G : Comonad C) : (Comonad.comparison G.adj).Full where map_surjective f := ⟨⟨f.f, by simpa using f.h⟩, rfl⟩ diff --git a/Mathlib/CategoryTheory/Monad/Kleisli.lean b/Mathlib/CategoryTheory/Monad/Kleisli.lean index e432473e5223b..9a251c5cb8b91 100644 --- a/Mathlib/CategoryTheory/Monad/Kleisli.lean +++ b/Mathlib/CategoryTheory/Monad/Kleisli.lean @@ -132,7 +132,7 @@ namespace Adjunction @[simps] def toCokleisli : C ⥤ Cokleisli U where obj X := (X : Cokleisli U) - map {X} {_} f := (U.ε.app X ≫ f : _) + map {X} {_} f := (U.ε.app X ≫ f :) map_comp {X} {Y} {_} f g := by -- Porting note: working around lack of unfold_projs change U.ε.app X ≫ f ≫ g = U.δ.app X ≫ U.map (U.ε.app X ≫ f) ≫ U.ε.app Y ≫ g diff --git a/Mathlib/CategoryTheory/Monad/Limits.lean b/Mathlib/CategoryTheory/Monad/Limits.lean index 8e99e8eace74d..6316344046718 100644 --- a/Mathlib/CategoryTheory/Monad/Limits.lean +++ b/Mathlib/CategoryTheory/Monad/Limits.lean @@ -347,7 +347,7 @@ theorem hasColimitsOfShape_of_reflective (R : D ⥤ C) [Reflective R] [HasColimi let t : IsColimit c := isColimitOfPreserves (monadicLeftAdjoint R) (colimit.isColimit _) apply HasColimit.mk ⟨_, (IsColimit.precomposeInvEquiv _ _).symm t⟩ apply - (isoWhiskerLeft F (asIso (monadicAdjunction R).counit) : _) ≪≫ F.rightUnitor + (isoWhiskerLeft F (asIso (monadicAdjunction R).counit) :) ≪≫ F.rightUnitor /-- If `C` has colimits then any reflective subcategory has colimits. -/ theorem hasColimits_of_reflective (R : D ⥤ C) [Reflective R] [HasColimitsOfSize.{v, u} C] : @@ -671,7 +671,7 @@ theorem hasLimitsOfShape_of_coreflective (R : D ⥤ C) [Coreflective R] [HasLimi let t : IsLimit c := isLimitOfPreserves (comonadicRightAdjoint R) (limit.isLimit _) apply HasLimit.mk ⟨_, (IsLimit.postcomposeHomEquiv _ _).symm t⟩ apply - (F.rightUnitor ≪≫ (isoWhiskerLeft F ((asIso (comonadicAdjunction R).unit) : _) )).symm + (F.rightUnitor ≪≫ (isoWhiskerLeft F ((asIso (comonadicAdjunction R).unit) :) )).symm /-- If `C` has limits then any coreflective subcategory has limits. -/ theorem hasLimits_of_coreflective (R : D ⥤ C) [Coreflective R] [HasLimitsOfSize.{v, u} C] : diff --git a/Mathlib/CategoryTheory/Monoidal/Braided/Basic.lean b/Mathlib/CategoryTheory/Monoidal/Braided/Basic.lean index b2025fc3fd6e4..4f1f7e9e08cb8 100644 --- a/Mathlib/CategoryTheory/Monoidal/Braided/Basic.lean +++ b/Mathlib/CategoryTheory/Monoidal/Braided/Basic.lean @@ -236,7 +236,7 @@ noncomputable def braidedCategoryOfFullyFaithful {C D : Type*} [Category C] [Cat braidedCategoryOfFaithful F (fun X Y => F.preimageIso ((μIso F _ _).symm ≪≫ β_ (F.obj X) (F.obj Y) ≪≫ (μIso F _ _))) - (by aesop_cat) + (by simp) section @@ -268,11 +268,9 @@ theorem braiding_leftUnitor_aux₂ (X : C) : ((β_ X (𝟙_ C)).hom ▷ 𝟙_ C) ≫ (α_ _ _ _).hom ≫ (α_ _ _ _).inv ≫ ((λ_ X).hom ▷ 𝟙_ C) := by monoidal _ = ((β_ X (𝟙_ C)).hom ▷ 𝟙_ C) ≫ (α_ _ _ _).hom ≫ (_ ◁ (β_ X _).hom) ≫ - (_ ◁ (β_ X _).inv) ≫ (α_ _ _ _).inv ≫ ((λ_ X).hom ▷ 𝟙_ C) := by - simp + (_ ◁ (β_ X _).inv) ≫ (α_ _ _ _).inv ≫ ((λ_ X).hom ▷ 𝟙_ C) := by simp _ = (α_ _ _ _).hom ≫ (β_ _ _).hom ≫ (α_ _ _ _).hom ≫ (_ ◁ (β_ X _).inv) ≫ (α_ _ _ _).inv ≫ - ((λ_ X).hom ▷ 𝟙_ C) := by - (slice_lhs 1 3 => rw [← hexagon_forward]); simp only [assoc] + ((λ_ X).hom ▷ 𝟙_ C) := by simp _ = (α_ _ _ _).hom ≫ (β_ _ _).hom ≫ ((λ_ _).hom ▷ X) ≫ (β_ X _).inv := by rw [braiding_leftUnitor_aux₁] _ = (α_ _ _ _).hom ≫ (_ ◁ (λ_ _).hom) ≫ (β_ _ _).hom ≫ (β_ X _).inv := by @@ -302,8 +300,7 @@ theorem braiding_rightUnitor_aux₂ (X : C) : _ = (α_ _ _ _).inv ≫ (β_ _ _).hom ≫ (α_ _ _ _).inv ≫ ((β_ _ X).inv ▷ _) ≫ (α_ _ _ _).hom ≫ (𝟙_ C ◁ (ρ_ X).hom) := by (slice_lhs 1 3 => rw [← hexagon_reverse]); simp only [assoc] - _ = (α_ _ _ _).inv ≫ (β_ _ _).hom ≫ (X ◁ (ρ_ _).hom) ≫ (β_ _ X).inv := by - rw [braiding_rightUnitor_aux₁] + _ = (α_ _ _ _).inv ≫ (β_ _ _).hom ≫ (X ◁ (ρ_ _).hom) ≫ (β_ _ X).inv := by simp _ = (α_ _ _ _).inv ≫ ((ρ_ _).hom ▷ _) ≫ (β_ _ X).hom ≫ (β_ _ _).inv := by (slice_lhs 2 3 => rw [← braiding_naturality_left]); simp only [assoc] _ = (α_ _ _ _).inv ≫ ((ρ_ _).hom ▷ _) := by rw [Iso.hom_inv_id, comp_id] diff --git a/Mathlib/CategoryTheory/Monoidal/End.lean b/Mathlib/CategoryTheory/Monoidal/End.lean index ab0ce6bcf4ec1..2994d7b5dc0ef 100644 --- a/Mathlib/CategoryTheory/Monoidal/End.lean +++ b/Mathlib/CategoryTheory/Monoidal/End.lean @@ -141,7 +141,7 @@ theorem ε_naturality {X Y : C} (f : X ⟶ Y) [F.LaxMonoidal] : @[reassoc (attr := simp)] theorem η_naturality {X Y : C} (f : X ⟶ Y) [F.OplaxMonoidal]: (η F).app X ≫ (𝟙_ (C ⥤ C)).map f = (η F).app X ≫ f := by - aesop_cat + simp @[reassoc (attr := simp)] theorem μ_naturality {m n : M} {X Y : C} (f : X ⟶ Y) [F.LaxMonoidal] : diff --git a/Mathlib/CategoryTheory/Monoidal/Internal/Limits.lean b/Mathlib/CategoryTheory/Monoidal/Internal/Limits.lean index 84cdce3c91ae3..5125b23e93c80 100644 --- a/Mathlib/CategoryTheory/Monoidal/Internal/Limits.lean +++ b/Mathlib/CategoryTheory/Monoidal/Internal/Limits.lean @@ -52,7 +52,7 @@ def limitCone (F : J ⥤ Mon_ C) : Cone F where -/ def forgetMapConeLimitConeIso (F : J ⥤ Mon_ C) : (forget C).mapCone (limitCone F) ≅ limit.cone (F ⋙ forget C) := - Cones.ext (Iso.refl _) (by aesop_cat) + Cones.ext (Iso.refl _) (by simp) /-- Implementation of `Mon_.hasLimitsOfShape`: the proposed cone over a functor `F : J ⥤ Mon_ C` is a limit cone. diff --git a/Mathlib/CategoryTheory/Monoidal/Subcategory.lean b/Mathlib/CategoryTheory/Monoidal/Subcategory.lean index 52293104297ba..407983573fdc0 100644 --- a/Mathlib/CategoryTheory/Monoidal/Subcategory.lean +++ b/Mathlib/CategoryTheory/Monoidal/Subcategory.lean @@ -128,7 +128,7 @@ instance fullBraidedSubcategory : BraidedCategory (FullSubcategory P) := braidedCategoryOfFaithful (fullSubcategoryInclusion P) (fun X Y => ⟨(β_ X.1 Y.1).hom, (β_ X.1 Y.1).inv, (β_ X.1 Y.1).hom_inv_id, (β_ X.1 Y.1).inv_hom_id⟩) - fun X Y => by aesop_cat + fun X Y => by simp /-- The forgetful braided functor from a full braided subcategory into the original category ("forgetting" the condition). diff --git a/Mathlib/CategoryTheory/NatIso.lean b/Mathlib/CategoryTheory/NatIso.lean index e12d02fb99a67..caa4309d21b37 100644 --- a/Mathlib/CategoryTheory/NatIso.lean +++ b/Mathlib/CategoryTheory/NatIso.lean @@ -227,7 +227,7 @@ theorem ofComponents.app (app' : ∀ X : C, F.obj X ≅ G.obj X) (naturality) (X /-- A natural transformation is an isomorphism if all its components are isomorphisms. -/ theorem isIso_of_isIso_app (α : F ⟶ G) [∀ X : C, IsIso (α.app X)] : IsIso α := - (ofComponents (fun X => asIso (α.app X)) (by aesop)).isIso_hom + (ofComponents (fun X => asIso (α.app X)) (by simp)).isIso_hom /-- Horizontal composition of natural isomorphisms. -/ @[simps] diff --git a/Mathlib/CategoryTheory/Opposites.lean b/Mathlib/CategoryTheory/Opposites.lean index 8cb8e3e027660..914d995a8d4cd 100644 --- a/Mathlib/CategoryTheory/Opposites.lean +++ b/Mathlib/CategoryTheory/Opposites.lean @@ -120,7 +120,7 @@ end /-- If `f` is an isomorphism, so is `f.op` -/ instance isIso_op {X Y : C} (f : X ⟶ Y) [IsIso f] : IsIso f.op := - ⟨⟨(inv f).op, ⟨Quiver.Hom.unop_inj (by aesop_cat), Quiver.Hom.unop_inj (by aesop_cat)⟩⟩⟩ + ⟨⟨(inv f).op, ⟨Quiver.Hom.unop_inj (by simp), Quiver.Hom.unop_inj (by simp)⟩⟩⟩ /-- If `f.op` is an isomorphism `f` must be too. (This cannot be an instance as it would immediately loop!) diff --git a/Mathlib/CategoryTheory/Pi/Basic.lean b/Mathlib/CategoryTheory/Pi/Basic.lean index 620eae9dd1af8..0b47b17e1e3a3 100644 --- a/Mathlib/CategoryTheory/Pi/Basic.lean +++ b/Mathlib/CategoryTheory/Pi/Basic.lean @@ -109,7 +109,7 @@ def comapComp (f : K → J) (g : J → I) : comap C g ⋙ comap (C ∘ g) f ≅ /-- The natural isomorphism between pulling back then evaluating, and just evaluating. -/ @[simps!] def comapEvalIsoEval (h : J → I) (j : J) : comap C h ⋙ eval (C ∘ h) j ≅ eval C (h j) := - NatIso.ofComponents (fun _ => Iso.refl _) (by simp only [Iso.refl]; aesop_cat) + NatIso.ofComponents (fun _ => Iso.refl _) (by simp only [Iso.refl]; simp) end diff --git a/Mathlib/CategoryTheory/Preadditive/Biproducts.lean b/Mathlib/CategoryTheory/Preadditive/Biproducts.lean index 49104d0c125a5..4a9c75efd68f9 100644 --- a/Mathlib/CategoryTheory/Preadditive/Biproducts.lean +++ b/Mathlib/CategoryTheory/Preadditive/Biproducts.lean @@ -143,7 +143,7 @@ def isBilimitOfIsLimit {f : J → C} (t : Bicone f) (ht : IsLimit t.toCone) : t. /-- We can turn any limit cone over a pair into a bilimit bicone. -/ def biconeIsBilimitOfLimitConeOfIsLimit {f : J → C} {t : Cone (Discrete.functor f)} (ht : IsLimit t) : (Bicone.ofLimitCone ht).IsBilimit := - isBilimitOfIsLimit _ <| IsLimit.ofIsoLimit ht <| Cones.ext (Iso.refl _) (by aesop_cat) + isBilimitOfIsLimit _ <| IsLimit.ofIsoLimit ht <| Cones.ext (Iso.refl _) (by simp) /-- In a preadditive category, any finite bicone which is a colimit cocone is in fact a bilimit bicone. -/ diff --git a/Mathlib/CategoryTheory/Preadditive/Mat.lean b/Mathlib/CategoryTheory/Preadditive/Mat.lean index af0bc591a2f18..4a2228de6b269 100644 --- a/Mathlib/CategoryTheory/Preadditive/Mat.lean +++ b/Mathlib/CategoryTheory/Preadditive/Mat.lean @@ -380,7 +380,7 @@ theorem additiveObjIsoBiproduct_naturality (F : Mat_ C ⥤ D) [Functor.Additive theorem additiveObjIsoBiproduct_naturality' (F : Mat_ C ⥤ D) [Functor.Additive F] {M N : Mat_ C} (f : M ⟶ N) : (additiveObjIsoBiproduct F M).inv ≫ F.map f = - biproduct.matrix (fun i j => F.map ((embedding C).map (f i j)) : _) ≫ + biproduct.matrix (fun i j => F.map ((embedding C).map (f i j)) :) ≫ (additiveObjIsoBiproduct F N).inv := by rw [Iso.inv_comp_eq, ← Category.assoc, Iso.eq_comp_inv, additiveObjIsoBiproduct_naturality] diff --git a/Mathlib/CategoryTheory/Preadditive/OfBiproducts.lean b/Mathlib/CategoryTheory/Preadditive/OfBiproducts.lean index c2f2f2b51b091..499eed7f4a959 100644 --- a/Mathlib/CategoryTheory/Preadditive/OfBiproducts.lean +++ b/Mathlib/CategoryTheory/Preadditive/OfBiproducts.lean @@ -51,12 +51,12 @@ theorem isUnital_leftAdd : EckmannHilton.IsUnital (· +ₗ ·) 0 := by have hr : ∀ f : X ⟶ Y, biprod.lift (0 : X ⟶ Y) f = f ≫ biprod.inr := by intro f ext - · aesop_cat + · simp · simp [biprod.lift_fst, Category.assoc, biprod.inr_fst, comp_zero] have hl : ∀ f : X ⟶ Y, biprod.lift f (0 : X ⟶ Y) = f ≫ biprod.inl := by intro f ext - · aesop_cat + · simp · simp [biprod.lift_snd, Category.assoc, biprod.inl_snd, comp_zero] exact { left_id := fun f => by simp [hr f, leftAdd, Category.assoc, Category.comp_id, biprod.inr_desc], @@ -67,12 +67,12 @@ theorem isUnital_rightAdd : EckmannHilton.IsUnital (· +ᵣ ·) 0 := by have h₂ : ∀ f : X ⟶ Y, biprod.desc (0 : X ⟶ Y) f = biprod.snd ≫ f := by intro f ext - · aesop_cat + · simp · simp only [biprod.inr_desc, BinaryBicone.inr_snd_assoc] have h₁ : ∀ f : X ⟶ Y, biprod.desc f (0 : X ⟶ Y) = biprod.fst ≫ f := by intro f ext - · aesop_cat + · simp · simp only [biprod.inr_desc, BinaryBicone.inr_fst_assoc, zero_comp] exact { left_id := fun f => by simp [h₂ f, rightAdd, biprod.lift_snd_assoc, Category.id_comp], diff --git a/Mathlib/CategoryTheory/Shift/Adjunction.lean b/Mathlib/CategoryTheory/Shift/Adjunction.lean index f893e59286f66..8f6a5a6ddc1f9 100644 --- a/Mathlib/CategoryTheory/Shift/Adjunction.lean +++ b/Mathlib/CategoryTheory/Shift/Adjunction.lean @@ -200,9 +200,31 @@ class CommShift : Prop where commShift_unit : NatTrans.CommShift adj.unit A := by infer_instance commShift_counit : NatTrans.CommShift adj.counit A := by infer_instance +open CommShift in +attribute [instance] commShift_unit commShift_counit + +@[reassoc (attr := simp)] +lemma unit_app_commShiftIso_hom_app [adj.CommShift A] (a : A) (X : C) : + adj.unit.app (X⟦a⟧) ≫ ((F ⋙ G).commShiftIso a).hom.app X = (adj.unit.app X)⟦a⟧' := by + simpa using (NatTrans.shift_app_comm adj.unit a X).symm + +@[reassoc (attr := simp)] +lemma unit_app_shift_commShiftIso_inv_app [adj.CommShift A] (a : A) (X : C) : + (adj.unit.app X)⟦a⟧' ≫ ((F ⋙ G).commShiftIso a).inv.app X = adj.unit.app (X⟦a⟧) := by + simp [← cancel_mono (((F ⋙ G).commShiftIso _).hom.app _)] + +@[reassoc (attr := simp)] +lemma commShiftIso_hom_app_counit_app_shift [adj.CommShift A] (a : A) (Y : D) : + ((G ⋙ F).commShiftIso a).hom.app Y ≫ (adj.counit.app Y)⟦a⟧' = adj.counit.app (Y⟦a⟧) := by + simpa using (NatTrans.shift_app_comm adj.counit a Y) + +@[reassoc (attr := simp)] +lemma commShiftIso_inv_app_counit_app [adj.CommShift A] (a : A) (Y : D) : + ((G ⋙ F).commShiftIso a).inv.app Y ≫ adj.counit.app (Y⟦a⟧) = (adj.counit.app Y)⟦a⟧' := by + simp [← cancel_epi (((G ⋙ F).commShiftIso _).hom.app _)] + namespace CommShift -attribute [instance] commShift_unit commShift_counit /-- Constructor for `Adjunction.CommShift`. -/ lemma mk' (h : NatTrans.CommShift adj.unit A) : @@ -217,6 +239,30 @@ lemma mk' (h : NatTrans.CommShift adj.unit A) : Functor.commShiftIso_id_hom_app, whiskerRight_app, id_comp, Functor.commShiftIso_comp_hom_app] using congr_app (h.shift_comm a) X⟩ +variable [adj.CommShift A] + +/-- The identity adjunction is compatible with the trivial `CommShift` structure on the +identity functor. +-/ +instance instId : (Adjunction.id (C := C)).CommShift A where + commShift_counit := + inferInstanceAs (NatTrans.CommShift (𝟭 C).leftUnitor.hom A) + commShift_unit := + inferInstanceAs (NatTrans.CommShift (𝟭 C).leftUnitor.inv A) + +variable {E : Type*} [Category E] {F' : D ⥤ E} {G' : E ⥤ D} (adj' : F' ⊣ G') + [HasShift E A] [F'.CommShift A] [G'.CommShift A] [adj.CommShift A] [adj'.CommShift A] + +/-- Compatibility of `Adjunction.Commshift` with the composition of adjunctions. +-/ +instance instComp : (adj.comp adj').CommShift A where + commShift_counit := by + rw [comp_counit] + infer_instance + commShift_unit := by + rw [comp_unit] + infer_instance + end CommShift variable {A} @@ -288,8 +334,8 @@ lemma iso_inv_app (Y : D) : (F.obj (G.obj Y)))) ≫ G.map ((shiftFunctor D a).map (adj.counit.app Y)) := by obtain rfl : b = -a := by rw [← add_left_inj a, h, neg_add_cancel] - simp only [Functor.comp_obj, iso, iso', shiftEquiv', Equiv.toFun_as_coe, - conjugateIsoEquiv_apply_inv, conjugateEquiv_apply_app, comp_unit_app, Functor.id_obj, + simp only [iso, iso', shiftEquiv', Equiv.toFun_as_coe, conjugateIsoEquiv_apply_inv, + conjugateEquiv_apply_app, Functor.comp_obj, comp_unit_app, Functor.id_obj, Equivalence.toAdjunction_unit, Equivalence.Equivalence_mk'_unit, Iso.symm_hom, Functor.comp_map, comp_counit_app, Equivalence.toAdjunction_counit, Equivalence.Equivalence_mk'_counit, Functor.map_shiftFunctorCompIsoId_hom_app, assoc, Functor.map_comp] @@ -470,9 +516,28 @@ lemma mk' (h : NatTrans.CommShift E.unitIso.hom A) : commShift_counit := (Adjunction.CommShift.mk' E.toAdjunction A h).commShift_counit /-- -If `E : C ≌ D` is an equivalence and we have compatible `CommShift` structures on `E.functor` -and `E.inverse`, then we also have compatible `CommShift` structures on `E.symm.functor` -and `E.symm.inverse`. +The forward functor of the identity equivalence is compatible with shifts. +-/ +instance : (Equivalence.refl (C := C)).functor.CommShift A := by + dsimp + infer_instance + +/-- +The inverse functor of the identity equivalence is compatible with shifts. +-/ +instance : (Equivalence.refl (C := C)).inverse.CommShift A := by + dsimp + infer_instance + +/-- +The identity equivalence is compatible with shifts. +-/ +instance : (Equivalence.refl (C := C)).CommShift A := by + dsimp [Equivalence.CommShift, refl_toAdjunction] + infer_instance + +/-- +If an equivalence `E : C ≌ D` is compatible with shifts, so is `E.symm`. -/ instance [E.CommShift A] : E.symm.CommShift A := mk' E.symm A (inferInstanceAs (NatTrans.CommShift E.counitIso.inv A)) @@ -483,6 +548,31 @@ lemma mk'' (h : NatTrans.CommShift E.counitIso.hom A) : have := mk' E.symm A (inferInstanceAs (NatTrans.CommShift E.counitIso.inv A)) inferInstanceAs (E.symm.symm.CommShift A) +variable {F : Type*} [Category F] [HasShift F A] {E' : D ≌ F} [E.CommShift A] + [E'.functor.CommShift A] [E'.inverse.CommShift A] [E'.CommShift A] + +/-- +If `E : C ≌ D` and `E' : D ≌ F` are equivalence whose forward functors are compatible with shifts, +so is `(E.trans E').functor`. +-/ +instance : (E.trans E').functor.CommShift A := by + dsimp + infer_instance + +/-- +If `E : C ≌ D` and `E' : D ≌ F` are equivalence whose inverse functors are compatible with shifts, +so is `(E.trans E').inverse`. +-/ +instance : (E.trans E').inverse.CommShift A := by + dsimp + infer_instance + +/-- +If equivalences `E : C ≌ D` and `E' : D ≌ F` are compatible with shifts, so is `E.trans E'`. +-/ +instance : (E.trans E').CommShift A := + inferInstanceAs ((E.toAdjunction.comp E'.toAdjunction).CommShift A) + end CommShift end diff --git a/Mathlib/CategoryTheory/Sigma/Basic.lean b/Mathlib/CategoryTheory/Sigma/Basic.lean index d7550a04b3288..c4076b46e5492 100644 --- a/Mathlib/CategoryTheory/Sigma/Basic.lean +++ b/Mathlib/CategoryTheory/Sigma/Basic.lean @@ -210,9 +210,9 @@ variable {I} {K : Type w₃} -- so that the suitable category instances could be found /-- The functor `Sigma.map` applied to a composition is a composition of functors. -/ @[simps!] -def mapComp (f : K → J) (g : J → I) : map (fun x ↦ C (g x)) f ⋙ (map C g : _) ≅ map C (g ∘ f) := +def mapComp (f : K → J) (g : J → I) : map (fun x ↦ C (g x)) f ⋙ (map C g :) ≅ map C (g ∘ f) := (descUniq _ _) fun k => - (isoWhiskerRight (inclCompMap (fun i => C (g i)) f k) (map C g : _) : _) ≪≫ inclCompMap _ _ _ + (isoWhiskerRight (inclCompMap (fun i => C (g i)) f k) (map C g :) :) ≪≫ inclCompMap _ _ _ end diff --git a/Mathlib/CategoryTheory/Simple.lean b/Mathlib/CategoryTheory/Simple.lean index 619a1bdfe7847..3d1d618bb6c68 100644 --- a/Mathlib/CategoryTheory/Simple.lean +++ b/Mathlib/CategoryTheory/Simple.lean @@ -116,7 +116,7 @@ variable (C) /-- We don't want the definition of 'simple' to include the zero object, so we check that here. -/ theorem zero_not_simple [Simple (0 : C)] : False := - (Simple.mono_isIso_iff_nonzero (0 : (0 : C) ⟶ (0 : C))).mp ⟨⟨0, by aesop_cat⟩⟩ rfl + (Simple.mono_isIso_iff_nonzero (0 : (0 : C) ⟶ (0 : C))).mp ⟨⟨0, by simp⟩⟩ rfl end diff --git a/Mathlib/CategoryTheory/SingleObj.lean b/Mathlib/CategoryTheory/SingleObj.lean index a3e3ca9146610..48fed64ab7d7d 100644 --- a/Mathlib/CategoryTheory/SingleObj.lean +++ b/Mathlib/CategoryTheory/SingleObj.lean @@ -205,11 +205,11 @@ def toSingleObjEquiv (e : M ≃* N) : SingleObj M ≌ SingleObj N where unitIso := eqToIso (by rw [← MonoidHom.comp_toFunctor, ← MonoidHom.id_toFunctor] congr 1 - aesop_cat) + simp) counitIso := eqToIso (by rw [← MonoidHom.comp_toFunctor, ← MonoidHom.id_toFunctor] congr 1 - aesop_cat) + simp) end MulEquiv diff --git a/Mathlib/CategoryTheory/Sites/DenseSubsite/Basic.lean b/Mathlib/CategoryTheory/Sites/DenseSubsite/Basic.lean index a7b6c780ed15f..e857f880260dd 100644 --- a/Mathlib/CategoryTheory/Sites/DenseSubsite/Basic.lean +++ b/Mathlib/CategoryTheory/Sites/DenseSubsite/Basic.lean @@ -172,11 +172,11 @@ theorem naturality [G.IsLocallyFull K] {X Y : C} (i : G.obj X ⟶ G.obj Y) : that is defined on a cover generated by the images of `G`. -/ noncomputable def pushforwardFamily {X} (x : ℱ.obj (op X)) : FamilyOfElements ℱ'.val (coverByImage G X) := fun _ _ hf => - ℱ'.val.map hf.some.lift.op <| α.app (op _) (ℱ.map hf.some.map.op x : _) + ℱ'.val.map hf.some.lift.op <| α.app (op _) (ℱ.map hf.some.map.op x) @[simp] theorem pushforwardFamily_def {X} (x : ℱ.obj (op X)) : pushforwardFamily α x = fun _ _ hf => - ℱ'.val.map hf.some.lift.op <| α.app (op _) (ℱ.map hf.some.map.op x : _) := rfl + ℱ'.val.map hf.some.lift.op <| α.app (op _) (ℱ.map hf.some.map.op x) := rfl @[simp] theorem pushforwardFamily_apply [G.IsLocallyFull K] diff --git a/Mathlib/CategoryTheory/Sites/EqualizerSheafCondition.lean b/Mathlib/CategoryTheory/Sites/EqualizerSheafCondition.lean index 2679d293f850b..29b141d4386aa 100644 --- a/Mathlib/CategoryTheory/Sites/EqualizerSheafCondition.lean +++ b/Mathlib/CategoryTheory/Sites/EqualizerSheafCondition.lean @@ -95,7 +95,7 @@ namespace Sieve to check a family is compatible. -/ def SecondObj : Type max v u := - ∏ᶜ fun f : Σ(Y Z : _) (_ : Z ⟶ Y), { f' : Y ⟶ X // S f' } => P.obj (op f.2.1) + ∏ᶜ fun f : Σ (Y Z : _) (_ : Z ⟶ Y), { f' : Y ⟶ X // S f' } => P.obj (op f.2.1) variable {P S} diff --git a/Mathlib/CategoryTheory/Sites/Pretopology.lean b/Mathlib/CategoryTheory/Sites/Pretopology.lean index d888a2611acf2..bde2d0fb7bd77 100644 --- a/Mathlib/CategoryTheory/Sites/Pretopology.lean +++ b/Mathlib/CategoryTheory/Sites/Pretopology.lean @@ -172,7 +172,7 @@ def trivial : Pretopology C where pullbacks X Y f S := by rintro ⟨Z, g, i, rfl⟩ refine ⟨pullback g f, pullback.snd _ _, ?_, ?_⟩ - · refine ⟨⟨pullback.lift (f ≫ inv g) (𝟙 _) (by simp), ⟨?_, by aesop_cat⟩⟩⟩ + · refine ⟨⟨pullback.lift (f ≫ inv g) (𝟙 _) (by simp), ⟨?_, by simp⟩⟩⟩ ext · rw [assoc, pullback.lift_fst, ← pullback.condition_assoc] simp diff --git a/Mathlib/CategoryTheory/Sites/Sheaf.lean b/Mathlib/CategoryTheory/Sites/Sheaf.lean index a94ed5480644d..04835ed2104f8 100644 --- a/Mathlib/CategoryTheory/Sites/Sheaf.lean +++ b/Mathlib/CategoryTheory/Sites/Sheaf.lean @@ -582,7 +582,6 @@ theorem isSheaf_iff_multiequalizer [∀ (X : C) (S : J.Cover X), HasMultiequaliz · apply (@asIso _ _ _ _ _ h).symm · intro a symm - erw [IsIso.inv_comp_eq] simp end MultiequalizerConditions diff --git a/Mathlib/CategoryTheory/Sites/Sheafification.lean b/Mathlib/CategoryTheory/Sites/Sheafification.lean index d8da476e5182c..be00c6666605c 100644 --- a/Mathlib/CategoryTheory/Sites/Sheafification.lean +++ b/Mathlib/CategoryTheory/Sites/Sheafification.lean @@ -136,7 +136,7 @@ variable {D} theorem isIso_toSheafify {P : Cᵒᵖ ⥤ D} (hP : Presheaf.IsSheaf J P) : IsIso (toSheafify J P) := by refine ⟨(sheafificationAdjunction J D |>.counit.app ⟨P, hP⟩).val, ?_, ?_⟩ - · change _ = (𝟙 (sheafToPresheaf J D ⋙ 𝟭 (Cᵒᵖ ⥤ D)) : _).app ⟨P, hP⟩ + · change _ = (𝟙 (sheafToPresheaf J D ⋙ 𝟭 (Cᵒᵖ ⥤ D)) :).app ⟨P, hP⟩ rw [← sheafificationAdjunction J D |>.right_triangle] rfl · change (sheafToPresheaf _ _).map _ ≫ _ = _ diff --git a/Mathlib/CategoryTheory/Sites/Subsheaf.lean b/Mathlib/CategoryTheory/Sites/Subsheaf.lean index be30cc3fa8ebc..b7b6807e53ccc 100644 --- a/Mathlib/CategoryTheory/Sites/Subsheaf.lean +++ b/Mathlib/CategoryTheory/Sites/Subsheaf.lean @@ -57,7 +57,7 @@ instance : PartialOrder (Subpresheaf F) := PartialOrder.lift Subpresheaf.obj (fun _ _ => Subpresheaf.ext) instance : Top (Subpresheaf F) := - ⟨⟨fun _ => ⊤, @fun U _ _ x _ => by aesop_cat⟩⟩ + ⟨⟨fun _ => ⊤, @fun U _ _ x _ => by simp⟩⟩ instance : Nonempty (Subpresheaf F) := inferInstance @@ -98,7 +98,7 @@ instance {G G' : Subpresheaf F} (h : G ≤ G') : Mono (Subpresheaf.homOfLe h) := NatTrans.ext <| funext fun U => funext fun x => - Subtype.ext <| (congr_arg Subtype.val <| (congr_fun (congr_app e U) x : _) : _)⟩ + Subtype.ext <| (congr_arg Subtype.val <| (congr_fun (congr_app e U) x :) :)⟩ @[reassoc (attr := simp)] theorem Subpresheaf.homOfLe_ι {G G' : Subpresheaf F} (h : G ≤ G') : @@ -194,7 +194,7 @@ theorem Subpresheaf.eq_sheafify (h : Presieve.IsSheaf J F) (hG : Presieve.IsShea exact ((hG _ hs).amalgamate _ (G.family_of_elements_compatible s)).2 apply (h _ hs).isSeparatedFor.ext intro V i hi - exact (congr_arg Subtype.val ((hG _ hs).valid_glue (G.family_of_elements_compatible s) _ hi) : _) + exact (congr_arg Subtype.val ((hG _ hs).valid_glue (G.family_of_elements_compatible s) _ hi) :) theorem Subpresheaf.sheafify_isSheaf (hF : Presieve.IsSheaf J F) : Presieve.IsSheaf J (G.sheafify J).toPresheaf := by @@ -293,7 +293,7 @@ theorem Subpresheaf.to_sheafify_lift_unique (h : Presieve.IsSheaf J F') rintro V i hi dsimp at hi rw [← FunctorToTypes.naturality, ← FunctorToTypes.naturality] - exact (congr_fun (congr_app e <| op V) ⟨_, hi⟩ : _) + exact (congr_fun (congr_app e <| op V) ⟨_, hi⟩ :) theorem Subpresheaf.sheafify_le (h : G ≤ G') (hF : Presieve.IsSheaf J F) (hG' : Presieve.IsSheaf J G'.toPresheaf) : G.sheafify J ≤ G' := by @@ -358,7 +358,7 @@ instance isIso_toImagePresheaf {F F' : Cᵒᵖ ⥤ (Type (max v w))} (f : F ⟶ · intro x y e have := (NatTrans.mono_iff_mono_app f).mp hf X rw [mono_iff_injective] at this - exact this (congr_arg Subtype.val e : _) + exact this (congr_arg Subtype.val e :) · rintro ⟨_, ⟨x, rfl⟩⟩ exact ⟨x, rfl⟩ apply NatIso.isIso_of_isIso_app diff --git a/Mathlib/CategoryTheory/Subobject/Comma.lean b/Mathlib/CategoryTheory/Subobject/Comma.lean index 20c1832042687..31c43bb091ef1 100644 --- a/Mathlib/CategoryTheory/Subobject/Comma.lean +++ b/Mathlib/CategoryTheory/Subobject/Comma.lean @@ -107,7 +107,7 @@ def subobjectEquiv [HasFiniteLimits C] [PreservesFiniteLimits T] (A : Structured refine ⟨fun h => Subobject.mk_le_mk_of_comm ?_ ?_, fun h => ?_⟩ · exact homMk (Subobject.ofMkLEMk _ _ h) ((cancel_mono (T.map g.right)).1 (by simp [← T.map_comp])) - · aesop_cat + · simp · refine Subobject.mk_le_mk_of_comm (Subobject.ofMkLEMk _ _ h).right ?_ exact congr_arg CommaMorphism.right (Subobject.ofMkLEMk_comp h) @@ -187,7 +187,7 @@ theorem lift_projectQuotient [HasFiniteColimits C] [PreservesFiniteColimits S] · exact (Subobject.underlyingIso f.unop.left.op).unop · refine (cancel_epi (S.map f.unop.left)).1 ?_ simpa [← Category.assoc, ← S.map_comp] using hq - · exact Quiver.Hom.unop_inj (by aesop_cat)) + · exact Quiver.Hom.unop_inj (by simp)) /-- Technical lemma for `quotientEquiv`. -/ theorem unop_left_comp_ofMkLEMk_unop {A : CostructuredArrow S T} {P Q : (CostructuredArrow S T)ᵒᵖ} diff --git a/Mathlib/CategoryTheory/Subobject/Lattice.lean b/Mathlib/CategoryTheory/Subobject/Lattice.lean index 813f8887842b6..70d9b061a97da 100644 --- a/Mathlib/CategoryTheory/Subobject/Lattice.lean +++ b/Mathlib/CategoryTheory/Subobject/Lattice.lean @@ -59,7 +59,7 @@ variable [HasPullbacks C] is (isomorphic to) the top object in `MonoOver X`. -/ def pullbackTop (f : X ⟶ Y) : (pullback f).obj ⊤ ≅ ⊤ := iso_of_both_ways (leTop _) - (homMk (pullback.lift f (𝟙 _) (by aesop_cat)) (pullback.lift_snd _ _ _)) + (homMk (pullback.lift f (𝟙 _) (by simp)) (pullback.lift_snd _ _ _)) /-- There is a morphism from `⊤ : MonoOver A` to the pullback of a monomorphism along itself; as the category is thin this is an isomorphism. -/ @@ -527,7 +527,7 @@ def leInfCone {A : C} (s : Set (Subobject A)) (f : Subobject A) (k : ∀ g ∈ s (by rcases j with ⟨-, ⟨g, ⟨m, rfl⟩⟩⟩ simpa using m)))) - (by aesop_cat) + (by simp) @[simp] theorem leInfCone_π_app_none {A : C} (s : Set (Subobject A)) (f : Subobject A) @@ -572,7 +572,7 @@ theorem sInf_le {A : C} (s : Set (Subobject A)) (f) (hf : f ∈ s) : sInf s ≤ simp only [Category.comp_id, Category.assoc, ← underlyingIso_hom_comp_eq_mk, Subobject.arrow_congr, congrArg_mpr_hom_left, Iso.cancel_iso_hom_left] convert limit.w (wideCospan s) (WidePullbackShape.Hom.term _) - aesop_cat + simp theorem le_sInf {A : C} (s : Set (Subobject A)) (f : Subobject A) (k : ∀ g ∈ s, f ≤ g) : f ≤ sInf s := by diff --git a/Mathlib/CategoryTheory/Triangulated/Adjunction.lean b/Mathlib/CategoryTheory/Triangulated/Adjunction.lean new file mode 100644 index 0000000000000..2308c4936e127 --- /dev/null +++ b/Mathlib/CategoryTheory/Triangulated/Adjunction.lean @@ -0,0 +1,204 @@ +/- +Copyright (c) 2024 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.Triangulated.Functor +import Mathlib.CategoryTheory.Shift.Adjunction +import Mathlib.CategoryTheory.Adjunction.Additive + +/-! +# The adjoint functor is triangulated + +If a functor `F : C ⥤ D` between pretriangulated categories is triangulated, and if we +have an adjunction `F ⊣ G`, then `G` is also a triangulated functor. + +We deduce that, if `E : C ≌ D` is an equivalence of pretriangulated categories, then +`E.functor` is triangulated if and only if `E.inverse` is triangulated. + +TODO: The case of left adjoints. +-/ + +namespace CategoryTheory + +open Category Limits Preadditive Pretriangulated Adjunction + +variable {C D : Type*} [Category C] [Category D] [HasZeroObject C] [HasZeroObject D] + [Preadditive C] [Preadditive D] [HasShift C ℤ] [HasShift D ℤ] + [∀ (n : ℤ), (shiftFunctor C n).Additive] [∀ (n : ℤ), (shiftFunctor D n).Additive] + [Pretriangulated C] [Pretriangulated D] + +namespace Adjunction + +variable {F : C ⥤ D} {G : D ⥤ C} (adj : F ⊣ G) [F.CommShift ℤ] [G.CommShift ℤ] + [adj.CommShift ℤ] + +include adj in +/-- +The right adjoint of a triangulated functor is triangulated. +-/ +lemma isTriangulated_rightAdjoint [F.IsTriangulated] : G.IsTriangulated where + map_distinguished T hT := by + have : G.Additive := adj.right_adjoint_additive + obtain ⟨Z, f, g, mem⟩ := distinguished_cocone_triangle (G.map T.mor₁) + obtain ⟨h, ⟨h₁, h₂⟩⟩ := complete_distinguished_triangle_morphism _ _ + (F.map_distinguished _ mem) hT (adj.counit.app T.obj₁) (adj.counit.app T.obj₂) (by simp) + dsimp at h h₁ h₂ ⊢ + have h₁' : f ≫ adj.unit.app Z ≫ G.map h = G.map T.mor₂ := by + simpa [homEquiv_apply] using DFunLike.congr_arg (adj.homEquiv _ _) h₁ + have h₂' : g ≫ (G.commShiftIso (1 : ℤ)).inv.app T.obj₁ = + adj.homEquiv _ _ h ≫ G.map T.mor₃ := by + apply (adj.homEquiv _ _).symm.injective + simp only [Functor.comp_obj, homEquiv_counit, Functor.id_obj, Functor.map_comp, assoc, + homEquiv_unit, counit_naturality, counit_naturality_assoc, left_triangle_components_assoc, + ← h₂, adj.shift_counit_app, Iso.hom_inv_id_app_assoc] + rw [assoc] at h₂ + have : Mono (adj.homEquiv _ _ h) := by + rw [mono_iff_cancel_zero] + intro _ φ hφ + obtain ⟨ψ, rfl⟩ := Triangle.coyoneda_exact₃ _ mem φ (by + dsimp + simp only [homEquiv_unit, Functor.comp_obj] at hφ + rw [← cancel_mono ((G.commShiftIso (1 : ℤ)).inv.app T.obj₁), assoc, h₂', zero_comp, + homEquiv_unit, assoc, reassoc_of% hφ, zero_comp]) + dsimp at ψ hφ ⊢ + obtain ⟨α, hα⟩ := T.coyoneda_exact₂ hT ((adj.homEquiv _ _).symm ψ) + ((adj.homEquiv _ _).injective (by simpa [homEquiv_counit, homEquiv_unit, ← h₁'] using hφ)) + have eq := DFunLike.congr_arg (adj.homEquiv _ _ ) hα + simp only [homEquiv_counit, Functor.id_obj, homEquiv_unit, comp_id, + Functor.map_comp, unit_naturality_assoc, right_triangle_components] at eq + have eq' := comp_distTriang_mor_zero₁₂ _ mem + dsimp at eq eq' + rw [eq, assoc, assoc, eq', comp_zero, comp_zero] + have := isIso_of_yoneda_map_bijective (adj.homEquiv _ _ h) (fun Y => by + constructor + · intro φ₁ φ₂ hφ + rw [← cancel_mono (adj.homEquiv _ _ h)] + exact hφ + · intro φ + obtain ⟨ψ, hψ⟩ := Triangle.coyoneda_exact₁ _ mem (φ ≫ G.map T.mor₃ ≫ + (G.commShiftIso (1 : ℤ)).hom.app T.obj₁) (by + dsimp + rw [assoc, assoc, ← G.commShiftIso_hom_naturality, ← G.map_comp_assoc, + comp_distTriang_mor_zero₃₁ _ hT, G.map_zero, zero_comp, comp_zero]) + dsimp at ψ hψ + obtain ⟨α, hα⟩ : ∃ α, α = φ - ψ ≫ adj.homEquiv _ _ h := ⟨_, rfl⟩ + have hα₀ : α ≫ G.map T.mor₃ = 0 := by + rw [hα, sub_comp, ← cancel_mono ((Functor.commShiftIso G (1 : ℤ)).hom.app T.obj₁), + assoc, sub_comp, assoc, assoc, hψ, zero_comp, sub_eq_zero, + ← cancel_mono ((Functor.commShiftIso G (1 : ℤ)).inv.app T.obj₁), assoc, + assoc, assoc, assoc, h₂', Iso.hom_inv_id_app, comp_id] + suffices ∃ (β : Y ⟶ Z), β ≫ adj.homEquiv _ _ h = α by + obtain ⟨β, hβ⟩ := this + refine ⟨ψ + β, ?_⟩ + dsimp + rw [add_comp, hβ, hα, add_sub_cancel] + obtain ⟨β, hβ⟩ := T.coyoneda_exact₃ hT ((adj.homEquiv _ _).symm α) + ((adj.homEquiv _ _).injective (by simpa [homEquiv_unit, homEquiv_counit] using hα₀)) + refine ⟨adj.homEquiv _ _ β ≫ f, ?_⟩ + simpa [homEquiv_unit, h₁'] using congr_arg (adj.homEquiv _ _).toFun hβ.symm) + refine isomorphic_distinguished _ mem _ (Iso.symm ?_) + refine Triangle.isoMk _ _ (Iso.refl _) (Iso.refl _) (asIso (adj.homEquiv Z T.obj₃ h)) ?_ ?_ ?_ + · dsimp + simp + · apply (adj.homEquiv _ _).symm.injective + dsimp + simp only [homEquiv_unit, homEquiv_counit, Functor.map_comp, assoc, + counit_naturality, left_triangle_components_assoc, h₁, id_comp] + · dsimp + rw [Functor.map_id, comp_id, homEquiv_unit, assoc, ← G.map_comp_assoc, ← h₂, + Functor.map_comp, Functor.map_comp, assoc, unit_naturality_assoc, assoc, + Functor.commShiftIso_hom_naturality, ← adj.shift_unit_app_assoc, + ← Functor.map_comp, right_triangle_components, Functor.map_id, comp_id] + +/-- +We say that an adjunction `F ⊣ G` is triangulated if it is compatible with the `CommShift` +structures on `F` and `G` (in the sense of `Adjunction.CommShift`) and if both `F` and `G` +are triangulated functors. +-/ +class IsTriangulated : Prop where + commShift : adj.CommShift ℤ := by infer_instance + leftAdjoint_isTriangulated : F.IsTriangulated := by infer_instance + rightAdjoint_isTriangulated : G.IsTriangulated := by infer_instance + +namespace IsTriangulated + +attribute [instance] commShift leftAdjoint_isTriangulated rightAdjoint_isTriangulated + +/-- Constructor for `Adjunction.IsTriangulated`. +-/ +lemma mk' [F.IsTriangulated] : adj.IsTriangulated where + rightAdjoint_isTriangulated := adj.isTriangulated_rightAdjoint + +/-- The identity adjunction is triangulated. +-/ +instance id : (Adjunction.id (C := C)).IsTriangulated where + +variable {E : Type*} [Category E] {F' : D ⥤ E} {G' : E ⥤ D} (adj' : F' ⊣ G') [HasZeroObject E] + [Preadditive E] [HasShift E ℤ] [∀ (n : ℤ), (shiftFunctor E n).Additive] [Pretriangulated E] + [F'.CommShift ℤ] [G'.CommShift ℤ] [adj'.CommShift ℤ] + +/-- A composition of triangulated adjunctions is triangulated. +-/ +instance comp [adj.IsTriangulated] [adj'.IsTriangulated] : (adj.comp adj').IsTriangulated where + +end IsTriangulated + +end Adjunction + +namespace Equivalence + +variable (E : C ≌ D) [E.functor.CommShift ℤ] [E.inverse.CommShift ℤ] [E.CommShift ℤ] + +/-- +We say that an equivalence of categories `E` is triangulated if both `E.functor` and +`E.inverse` are triangulated functors. +-/ +abbrev IsTriangulated : Prop := E.toAdjunction.IsTriangulated + +namespace IsTriangulated + +instance [E.IsTriangulated] : E.functor.IsTriangulated := inferInstance +instance [E.IsTriangulated] : E.inverse.IsTriangulated := inferInstance + +instance [h : E.functor.IsTriangulated] : E.symm.inverse.IsTriangulated := h +instance [h : E.inverse.IsTriangulated] : E.symm.functor.IsTriangulated := h + + +/-- Constructor for `Equivalence.IsTriangulated`. -/ +lemma mk' (h : E.functor.IsTriangulated) : E.IsTriangulated where + rightAdjoint_isTriangulated := E.toAdjunction.isTriangulated_rightAdjoint + +/-- Constructor for `Equivalence.IsTriangulated`. -/ +lemma mk'' (h : E.inverse.IsTriangulated) : E.IsTriangulated where + leftAdjoint_isTriangulated := (mk' E.symm h).rightAdjoint_isTriangulated + +/-- +The identity equivalence is triangulated. +-/ +instance refl : (Equivalence.refl (C := C)).IsTriangulated := by + dsimp [Equivalence.IsTriangulated] + rw [refl_toAdjunction] + infer_instance + +/-- If the equivalence `E` is triangulated, so is the equivalence `E.symm`. +-/ +instance symm [E.IsTriangulated] : E.symm.IsTriangulated where + +variable {D' : Type*} [Category D'] [HasZeroObject D'] [Preadditive D'] [HasShift D' ℤ] + [∀ (n : ℤ), (shiftFunctor D' n).Additive] [Pretriangulated D'] {E' : D ≌ D'} + [E'.functor.CommShift ℤ] [E'.inverse.CommShift ℤ] [E'.CommShift ℤ] + +/-- +If equivalences `E : C ≌ D` and `E' : D ≌ F` are triangulated, so is `E.trans E'`. +-/ +instance trans [E.IsTriangulated] [E'.IsTriangulated] : (E.trans E').IsTriangulated := by + dsimp [Equivalence.IsTriangulated] + rw [trans_toAdjunction] + infer_instance + +end IsTriangulated + +end Equivalence + +end CategoryTheory diff --git a/Mathlib/CategoryTheory/Triangulated/Basic.lean b/Mathlib/CategoryTheory/Triangulated/Basic.lean index e76c664d5f1d3..b8e3b05949917 100644 --- a/Mathlib/CategoryTheory/Triangulated/Basic.lean +++ b/Mathlib/CategoryTheory/Triangulated/Basic.lean @@ -249,7 +249,7 @@ def binaryProductTriangleIsoBinaryBiproductTriangle (X₁ X₂ : C) [HasZeroMorphisms C] [HasBinaryBiproduct X₁ X₂] : binaryProductTriangle X₁ X₂ ≅ binaryBiproductTriangle X₁ X₂ := Triangle.isoMk _ _ (Iso.refl _) (biprod.isoProd X₁ X₂).symm (Iso.refl _) - (by aesop_cat) (by aesop_cat) (by aesop_cat) + (by aesop_cat) (by simp) (by simp) section diff --git a/Mathlib/CategoryTheory/Triangulated/Functor.lean b/Mathlib/CategoryTheory/Triangulated/Functor.lean index c7c860c35beec..faa14d9b0ea07 100644 --- a/Mathlib/CategoryTheory/Triangulated/Functor.lean +++ b/Mathlib/CategoryTheory/Triangulated/Functor.lean @@ -76,7 +76,7 @@ noncomputable def mapTriangleCommShiftIso (n : ℤ) : Triangle.shiftFunctor C n ⋙ F.mapTriangle ≅ F.mapTriangle ⋙ Triangle.shiftFunctor D n := NatIso.ofComponents (fun T => Triangle.isoMk _ _ ((F.commShiftIso n).app _) ((F.commShiftIso n).app _) ((F.commShiftIso n).app _) - (by aesop_cat) (by aesop_cat) (by + (by simp) (by simp) (by dsimp simp only [map_units_smul, map_comp, Linear.units_smul_comp, assoc, Linear.comp_units_smul, ← F.commShiftIso_hom_naturality_assoc] @@ -103,7 +103,7 @@ def mapTriangleRotateIso : NatIso.ofComponents (fun T => Triangle.isoMk _ _ (Iso.refl _) (Iso.refl _) ((F.commShiftIso (1 : ℤ)).symm.app _) - (by aesop_cat) (by aesop_cat) (by aesop_cat)) (by aesop_cat) + (by simp) (by simp) (by simp)) (by aesop_cat) /-- `F.mapTriangle` commutes with the inverse of the rotation of triangles. -/ @[simps!] @@ -112,7 +112,7 @@ noncomputable def mapTriangleInvRotateIso [F.Additive] : Pretriangulated.invRotate C ⋙ F.mapTriangle := NatIso.ofComponents (fun T => Triangle.isoMk _ _ ((F.commShiftIso (-1 : ℤ)).symm.app _) (Iso.refl _) (Iso.refl _) - (by aesop_cat) (by aesop_cat) (by aesop_cat)) (by aesop_cat) + (by simp) (by simp) (by simp)) (by aesop_cat) variable (C) in diff --git a/Mathlib/CategoryTheory/Triangulated/Opposite/Pretriangulated.lean b/Mathlib/CategoryTheory/Triangulated/Opposite/Pretriangulated.lean index 5e146b7317e28..dcec2c68e78ac 100644 --- a/Mathlib/CategoryTheory/Triangulated/Opposite/Pretriangulated.lean +++ b/Mathlib/CategoryTheory/Triangulated/Opposite/Pretriangulated.lean @@ -88,7 +88,7 @@ noncomputable def contractibleTriangleIso (X : Cᵒᵖ) : rw [IsZero.iff_id_eq_zero] change (𝟙 ((0 : C)⟦(-1 : ℤ)⟧)).op = 0 rw [← Functor.map_id, id_zero, Functor.map_zero, op_zero])) - (by aesop_cat) (by aesop_cat) (by aesop_cat) + (by simp) (by simp) (by simp) lemma contractible_distinguished (X : Cᵒᵖ) : contractibleTriangle X ∈ distinguishedTriangles C := by @@ -103,7 +103,7 @@ noncomputable def rotateTriangleOpEquivalenceInverseObjRotateUnopIso (T : Triang ((triangleOpEquivalence C).inverse.obj T).unop := Triangle.isoMk _ _ (Iso.refl _) (Iso.refl _) (-((opShiftFunctorEquivalence C 1).unitIso.app T.obj₁).unop) (by simp) - (Quiver.Hom.op_inj (by aesop_cat)) (by aesop_cat) + (Quiver.Hom.op_inj (by simp)) (by simp) lemma rotate_distinguished_triangle (T : Triangle Cᵒᵖ) : T ∈ distinguishedTriangles C ↔ T.rotate ∈ distinguishedTriangles C := by @@ -119,7 +119,7 @@ lemma distinguished_cocone_triangle {X Y : Cᵒᵖ} (f : X ⟶ Y) : (shiftFunctor Cᵒᵖ (1 : ℤ)).map h.op, ?_⟩ simp only [mem_distinguishedTriangles_iff] refine Pretriangulated.isomorphic_distinguished _ H _ ?_ - exact Triangle.isoMk _ _ (Iso.refl _) (Iso.refl _) (Iso.refl _) (by aesop_cat) (by aesop_cat) + exact Triangle.isoMk _ _ (Iso.refl _) (Iso.refl _) (Iso.refl _) (by simp) (by simp) (Quiver.Hom.op_inj (by simp [shift_unop_opShiftFunctorEquivalence_counitIso_inv_app])) lemma complete_distinguished_triangle_morphism (T₁ T₂ : Triangle Cᵒᵖ) diff --git a/Mathlib/CategoryTheory/Triangulated/Opposite/Triangle.lean b/Mathlib/CategoryTheory/Triangulated/Opposite/Triangle.lean index cd0a07b4be3dd..2618748d4f279 100644 --- a/Mathlib/CategoryTheory/Triangulated/Opposite/Triangle.lean +++ b/Mathlib/CategoryTheory/Triangulated/Opposite/Triangle.lean @@ -73,7 +73,7 @@ equivalence `triangleOpEquivalence C : (Triangle C)ᵒᵖ ≌ Triangle Cᵒᵖ` @[simps!] noncomputable def unitIso : 𝟭 _ ≅ functor C ⋙ inverse C := NatIso.ofComponents (fun T => Iso.op - (Triangle.isoMk _ _ (Iso.refl _) (Iso.refl _) (Iso.refl _) (by aesop_cat) (by aesop_cat) + (Triangle.isoMk _ _ (Iso.refl _) (Iso.refl _) (Iso.refl _) (by simp) (by simp) (Quiver.Hom.op_inj (by simp [shift_unop_opShiftFunctorEquivalence_counitIso_inv_app])))) (fun {T₁ T₂} f => Quiver.Hom.unop_inj (by aesop_cat)) @@ -84,8 +84,8 @@ equivalence `triangleOpEquivalence C : (Triangle C)ᵒᵖ ≌ Triangle Cᵒᵖ` noncomputable def counitIso : inverse C ⋙ functor C ≅ 𝟭 _ := NatIso.ofComponents (fun T => by refine Triangle.isoMk _ _ (Iso.refl _) (Iso.refl _) (Iso.refl _) ?_ ?_ ?_ - · aesop_cat - · aesop_cat + · simp + · simp · dsimp rw [Functor.map_id, comp_id, id_comp, Functor.map_comp, ← opShiftFunctorEquivalence_counitIso_inv_naturality_assoc, diff --git a/Mathlib/CategoryTheory/Triangulated/Pretriangulated.lean b/Mathlib/CategoryTheory/Triangulated/Pretriangulated.lean index c3b954b379632..1cbad40237ea8 100644 --- a/Mathlib/CategoryTheory/Triangulated/Pretriangulated.lean +++ b/Mathlib/CategoryTheory/Triangulated/Pretriangulated.lean @@ -219,7 +219,7 @@ lemma contractible_distinguished₁ (X : C) : refine isomorphic_distinguished _ (inv_rot_of_distTriang _ (contractible_distinguished X)) _ ?_ exact Triangle.isoMk _ _ (Functor.mapZeroObject _).symm (Iso.refl _) (Iso.refl _) - (by aesop_cat) (by aesop_cat) (by aesop_cat) + (by simp) (by simp) (by simp) /-- Obvious triangles `X ⟶ 0 ⟶ X⟦1⟧ ⟶ X⟦1⟧` are distinguished -/ lemma contractible_distinguished₂ (X : C) : @@ -227,7 +227,7 @@ lemma contractible_distinguished₂ (X : C) : refine isomorphic_distinguished _ (inv_rot_of_distTriang _ (contractible_distinguished₁ (X⟦(1 : ℤ)⟧))) _ ?_ exact Triangle.isoMk _ _ ((shiftEquiv C (1 : ℤ)).unitIso.app X) (Iso.refl _) (Iso.refl _) - (by aesop_cat) (by aesop_cat) + (by simp) (by simp) (by dsimp; simp only [shift_shiftFunctorCompIsoId_inv_app, id_comp]) namespace Triangle @@ -531,7 +531,7 @@ lemma binaryBiproductTriangle_distinguished (X₁ X₂ : C) : dsimp at he₁ he₂ refine isomorphic_distinguished _ mem _ (Iso.symm ?_) refine Triangle.isoMk _ _ (Iso.refl _) e (Iso.refl _) - (by aesop_cat) (by aesop_cat) (by aesop_cat) + (by aesop_cat) (by aesop_cat) (by simp) lemma binaryProductTriangle_distinguished (X₁ X₂ : C) : binaryProductTriangle X₁ X₂ ∈ distTriang C := diff --git a/Mathlib/CategoryTheory/Triangulated/Rotate.lean b/Mathlib/CategoryTheory/Triangulated/Rotate.lean index 4ef56167fa648..92a42f9b69e8f 100644 --- a/Mathlib/CategoryTheory/Triangulated/Rotate.lean +++ b/Mathlib/CategoryTheory/Triangulated/Rotate.lean @@ -101,13 +101,9 @@ def invRotate : Triangle C ⥤ Triangle C where hom₃ := f.hom₂ comm₁ := by dsimp - simp only [neg_comp, assoc, comp_neg, neg_inj, ← Functor.map_comp_assoc, ← f.comm₃] - rw [Functor.map_comp, assoc] - erw [← NatTrans.naturality] - rfl - comm₃ := by - erw [← reassoc_of% f.comm₂, Category.assoc, ← NatTrans.naturality] - rfl } + simp only [comp_neg, ← Functor.map_comp_assoc, ← f.comm₃] + rw [Functor.map_comp] + simp } variable {C} variable [∀ n : ℤ, Functor.Additive (shiftFunctor C n)] diff --git a/Mathlib/CategoryTheory/Triangulated/Subcategory.lean b/Mathlib/CategoryTheory/Triangulated/Subcategory.lean index 22baf9e9ef9b8..f662bd84a092a 100644 --- a/Mathlib/CategoryTheory/Triangulated/Subcategory.lean +++ b/Mathlib/CategoryTheory/Triangulated/Subcategory.lean @@ -81,7 +81,7 @@ def isoClosure : Subcategory C where exact le_isoClosure _ _ (S.ext₂' (Triangle.mk (e₁.inv ≫ T.mor₁) (T.mor₂ ≫ e₃.hom) (e₃.inv ≫ T.mor₃ ≫ e₁.hom⟦1⟧')) (isomorphic_distinguished _ hT _ - (Triangle.isoMk _ _ e₁.symm (Iso.refl _) e₃.symm (by aesop_cat) (by aesop_cat) (by + (Triangle.isoMk _ _ e₁.symm (Iso.refl _) e₃.symm (by simp) (by simp) (by dsimp simp only [assoc, Iso.cancel_iso_inv_left, ← Functor.map_comp, e₁.hom_inv_id, Functor.map_id, comp_id]))) h₁ h₃) @@ -155,7 +155,7 @@ instance respectsIso_W : S.W.RespectsIso where precomp {X' X Y} e (he : IsIso e) := by rintro f ⟨Z, g, h, mem, mem'⟩ refine ⟨Z, g, h ≫ inv e⟦(1 : ℤ)⟧', isomorphic_distinguished _ mem _ ?_, mem'⟩ - refine Triangle.isoMk _ _ (asIso e) (Iso.refl _) (Iso.refl _) (by aesop_cat) (by aesop_cat) ?_ + refine Triangle.isoMk _ _ (asIso e) (Iso.refl _) (Iso.refl _) (by simp) (by simp) ?_ dsimp simp only [Functor.map_inv, assoc, IsIso.inv_hom_id, comp_id, id_comp] postcomp {X Y Y'} e (he : IsIso e) := by diff --git a/Mathlib/CategoryTheory/Triangulated/TriangleShift.lean b/Mathlib/CategoryTheory/Triangulated/TriangleShift.lean index b811588238be2..c06cd41d963ae 100644 --- a/Mathlib/CategoryTheory/Triangulated/TriangleShift.lean +++ b/Mathlib/CategoryTheory/Triangulated/TriangleShift.lean @@ -65,7 +65,7 @@ noncomputable def Triangle.shiftFunctorZero : Triangle.shiftFunctor C 0 ≅ 𝟭 NatIso.ofComponents (fun T => Triangle.isoMk _ _ ((CategoryTheory.shiftFunctorZero C ℤ).app _) ((CategoryTheory.shiftFunctorZero C ℤ).app _) ((CategoryTheory.shiftFunctorZero C ℤ).app _) - (by aesop_cat) (by aesop_cat) (by + (by simp) (by simp) (by dsimp simp only [one_smul, assoc, shiftFunctorComm_zero_hom_app, ← Functor.map_comp, Iso.inv_hom_id_app, Functor.id_obj, Functor.map_id, @@ -109,7 +109,7 @@ noncomputable def rotateRotateRotateIso : rotate C ⋙ rotate C ⋙ rotate C ≅ Triangle.shiftFunctor C 1 := NatIso.ofComponents (fun T => Triangle.isoMk _ _ (Iso.refl _) (Iso.refl _) (Iso.refl _) - (by aesop_cat) (by aesop_cat) (by aesop_cat)) + (by simp) (by simp) (by simp)) (by aesop_cat) /-- Rotating triangles three times backwards identifies with the shift by `-1`. -/ @@ -117,8 +117,8 @@ noncomputable def invRotateInvRotateInvRotateIso : invRotate C ⋙ invRotate C ⋙ invRotate C ≅ Triangle.shiftFunctor C (-1) := NatIso.ofComponents (fun T => Triangle.isoMk _ _ (Iso.refl _) (Iso.refl _) (Iso.refl _) - (by aesop_cat) - (by aesop_cat) + (by simp) + (by simp) (by dsimp [shiftFunctorCompIsoId] simp [shiftFunctorComm_eq C _ _ _ (add_neg_cancel (1 : ℤ))])) diff --git a/Mathlib/CategoryTheory/Yoneda.lean b/Mathlib/CategoryTheory/Yoneda.lean index 05d0c28efabb3..e362440ade3cf 100644 --- a/Mathlib/CategoryTheory/Yoneda.lean +++ b/Mathlib/CategoryTheory/Yoneda.lean @@ -523,7 +523,7 @@ def yonedaCompUliftFunctorEquiv (F : Cᵒᵖ ⥤ Type max v₁ w) (X : C) : dsimp rw [Category.comp_id] rfl - right_inv f := by aesop_cat + right_inv f := by simp /-- The Yoneda lemma asserts that the Yoneda pairing `(X : Cᵒᵖ, F : Cᵒᵖ ⥤ Type) ↦ (yoneda.obj (unop X) ⟶ F)` @@ -705,7 +705,7 @@ def coyonedaCompUliftFunctorEquiv (F : C ⥤ Type max v₁ w) (X : Cᵒᵖ) : dsimp rw [Category.id_comp] rfl - right_inv f := by aesop_cat + right_inv f := by simp /-- The Coyoneda lemma asserts that the Coyoneda pairing `(X : C, F : C ⥤ Type) ↦ (coyoneda.obj X ⟶ F)` diff --git a/Mathlib/Combinatorics/SetFamily/LYM.lean b/Mathlib/Combinatorics/SetFamily/LYM.lean index 67a603abd4165..24ec53fa6da54 100644 --- a/Mathlib/Combinatorics/SetFamily/LYM.lean +++ b/Mathlib/Combinatorics/SetFamily/LYM.lean @@ -211,7 +211,7 @@ end LYM /-- **Sperner's theorem**. The size of an antichain in `Finset α` is bounded by the size of the maximal layer in `Finset α`. This precisely means that `Finset α` is a Sperner order. -/ -theorem IsAntichain.sperner [Fintype α] {𝒜 : Finset (Finset α)} +theorem _root_.IsAntichain.sperner [Fintype α] {𝒜 : Finset (Finset α)} (h𝒜 : IsAntichain (· ⊆ ·) (𝒜 : Set (Finset α))) : #𝒜 ≤ (Fintype.card α).choose (Fintype.card α / 2) := by classical diff --git a/Mathlib/Combinatorics/SimpleGraph/Acyclic.lean b/Mathlib/Combinatorics/SimpleGraph/Acyclic.lean index 9782814a525cc..126e91256aa32 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Acyclic.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Acyclic.lean @@ -175,7 +175,7 @@ lemma IsTree.card_edgeFinset [Fintype V] [Fintype G.edgeSet] (hG : G.IsTree) : length_tail_add_one (not_nil_of_ne (by simpa using ha))] at h3 omega · simp only [ne_eq, eq_mp_eq_cast, id_eq, isPath_copy] - exact (hf _).tail (not_nil_of_ne (by simpa using ha)) + exact (hf _).tail case surj => simp only [mem_edgeFinset, Finset.mem_compl, Finset.mem_singleton, Sym2.forall, mem_edgeSet] intros x y h diff --git a/Mathlib/Combinatorics/SimpleGraph/Basic.lean b/Mathlib/Combinatorics/SimpleGraph/Basic.lean index fbf6473d0fddf..0c0f567f36892 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Basic.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Basic.lean @@ -197,6 +197,12 @@ theorem adj_injective : Injective (Adj : SimpleGraph V → V → V → Prop) := theorem adj_inj {G H : SimpleGraph V} : G.Adj = H.Adj ↔ G = H := adj_injective.eq_iff +theorem adj_congr_of_sym2 {u v w x : V} (h : s(u, v) = s(w, x)) : G.Adj u v ↔ G.Adj w x := by + simp only [Sym2.eq, Sym2.rel_iff', Prod.mk.injEq, Prod.swap_prod_mk] at h + cases' h with hl hr + · rw [hl.1, hl.2] + · rw [hr.1, hr.2, adj_comm] + section Order /-- The relation that one `SimpleGraph` is a subgraph of another. diff --git a/Mathlib/Combinatorics/SimpleGraph/Coloring.lean b/Mathlib/Combinatorics/SimpleGraph/Coloring.lean index 9d98315f91068..7f5fe6d81b937 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Coloring.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Coloring.lean @@ -265,10 +265,6 @@ theorem chromaticNumber_le_iff_colorable {n : ℕ} : G.chromaticNumber ≤ n ↔ rw [Set.mem_setOf_eq] at this exact this.mono h -@[deprecated Colorable.chromaticNumber_le (since := "2024-03-21")] -theorem chromaticNumber_le_card [Fintype α] (C : G.Coloring α) : - G.chromaticNumber ≤ Fintype.card α := C.colorable.chromaticNumber_le - theorem colorable_chromaticNumber {m : ℕ} (hc : G.Colorable m) : G.Colorable (ENat.toNat G.chromaticNumber) := by classical diff --git a/Mathlib/Combinatorics/SimpleGraph/Connectivity/WalkCounting.lean b/Mathlib/Combinatorics/SimpleGraph/Connectivity/WalkCounting.lean index 01d7e6a4232f3..10adfe87aff11 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Connectivity/WalkCounting.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Connectivity/WalkCounting.lean @@ -63,7 +63,7 @@ Note that `u` and `v` may be the same. -/ @[simps] def walkLengthTwoEquivCommonNeighbors (u v : V) : {p : G.Walk u v // p.length = 2} ≃ G.commonNeighbors u v where - toFun p := ⟨p.val.getVert 1, match p with + toFun p := ⟨p.val.snd, match p with | ⟨.cons _ (.cons _ .nil), _⟩ => ⟨‹G.Adj u _›, ‹G.Adj _ v›.symm⟩⟩ invFun w := ⟨w.prop.1.toWalk.concat w.prop.2.symm, rfl⟩ left_inv | ⟨.cons _ (.cons _ .nil), hp⟩ => by rfl diff --git a/Mathlib/Combinatorics/SimpleGraph/Operations.lean b/Mathlib/Combinatorics/SimpleGraph/Operations.lean index fea650ca66f09..4af95c2d77aa8 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Operations.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Operations.lean @@ -5,6 +5,7 @@ Authors: Jeremy Tan -/ import Mathlib.Combinatorics.SimpleGraph.Finite import Mathlib.Combinatorics.SimpleGraph.Maps +import Mathlib.Combinatorics.SimpleGraph.Subgraph /-! # Local graph operations @@ -168,6 +169,15 @@ lemma sup_edge_of_adj (h : G.Adj s t) : G ⊔ edge s t = G := by rwa [sup_eq_left, ← edgeSet_subset_edgeSet, edge_edgeSet_of_ne h.ne, Set.singleton_subset_iff, mem_edgeSet] +theorem Subgraph.spanningCoe_sup_edge_le {H : Subgraph (G ⊔ edge s t)} (h : ¬ H.Adj s t) : + H.spanningCoe ≤ G := by + intro v w hvw + have := hvw.adj_sub + simp only [Subgraph.spanningCoe_adj, SimpleGraph.sup_adj, SimpleGraph.edge_adj] at * + by_cases hs : s(v, w) = s(s, t) + · exact (h ((Subgraph.adj_congr_of_sym2 hs).mp hvw)).elim + · aesop + variable [Fintype V] [DecidableRel G.Adj] variable [DecidableEq V] in diff --git a/Mathlib/Combinatorics/SimpleGraph/Path.lean b/Mathlib/Combinatorics/SimpleGraph/Path.lean index 34fc605b7283d..25c395c316dff 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Path.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Path.lean @@ -258,10 +258,11 @@ lemma isCycle_reverse {p : G.Walk u u} : p.reverse.IsCycle ↔ p.IsCycle where mp h := by simpa using h.reverse mpr := .reverse -lemma IsPath.tail {p : G.Walk u v} (hp : p.IsPath) (hp' : ¬ p.Nil) : p.tail.IsPath := by - rw [Walk.isPath_def] at hp ⊢ - rw [← cons_support_tail _ hp', List.nodup_cons] at hp - exact hp.2 +lemma IsPath.tail {p : G.Walk u v} (hp : p.IsPath) : p.tail.IsPath := by + cases p with + | nil => simp + | cons hadj p => + simp_all [Walk.isPath_def] /-! ### About paths -/ @@ -274,6 +275,26 @@ theorem IsPath.length_lt [Fintype V] {u v : V} {p : G.Walk u v} (hp : p.IsPath) rw [Nat.lt_iff_add_one_le, ← length_support] exact hp.support_nodup.length_le_card +lemma IsPath.getVert_injOn {p : G.Walk u v} (hp : p.IsPath) : + Set.InjOn p.getVert {i | i ≤ p.length} := by + intro n hn m hm hnm + induction p generalizing n m with + | nil => aesop + | @cons v w u h p ihp => + simp only [length_cons, Set.mem_setOf_eq] at hn hm hnm + by_cases hn0 : n = 0 <;> by_cases hm0 : m = 0 + · aesop + · simp only [hn0, getVert_zero, Walk.getVert_cons p h hm0] at hnm + have hvp : v ∉ p.support := by aesop + exact (hvp (Walk.mem_support_iff_exists_getVert.mpr ⟨(m - 1), ⟨hnm.symm, by omega⟩⟩)).elim + · simp only [hm0, Walk.getVert_cons p h hn0] at hnm + have hvp : v ∉ p.support := by aesop + exact (hvp (Walk.mem_support_iff_exists_getVert.mpr ⟨(n - 1), ⟨hnm, by omega⟩⟩)).elim + · simp only [Walk.getVert_cons _ _ hn0, Walk.getVert_cons _ _ hm0] at hnm + have := ihp hp.of_cons (by omega : (n - 1) ≤ p.length) + (by omega : (m - 1) ≤ p.length) hnm + omega + /-! ### Walk decompositions -/ section WalkDecomp diff --git a/Mathlib/Combinatorics/SimpleGraph/Subgraph.lean b/Mathlib/Combinatorics/SimpleGraph/Subgraph.lean index 3ee260f7663ee..87264b888eb6b 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Subgraph.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Subgraph.lean @@ -118,6 +118,13 @@ protected theorem Adj.snd_mem {H : G.Subgraph} {u v : V} (h : H.Adj u v) : v ∈ protected theorem Adj.ne {H : G.Subgraph} {u v : V} (h : H.Adj u v) : u ≠ v := h.adj_sub.ne +theorem adj_congr_of_sym2 {H : G.Subgraph} {u v w x : V} (h2 : s(u, v) = s(w, x)) : + H.Adj u v ↔ H.Adj w x := by + simp only [Sym2.eq, Sym2.rel_iff', Prod.mk.injEq, Prod.swap_prod_mk] at h2 + cases' h2 with hl hr + · rw [hl.1, hl.2] + · rw [hr.1, hr.2, Subgraph.adj_comm] + /-- Coercion from `G' : Subgraph G` to a `SimpleGraph G'.verts`. -/ @[simps] protected def coe (G' : Subgraph G) : SimpleGraph G'.verts where @@ -164,6 +171,9 @@ lemma spanningCoe_le (G' : G.Subgraph) : G'.spanningCoe ≤ G := fun _ _ ↦ G'. theorem spanningCoe_inj : G₁.spanningCoe = G₂.spanningCoe ↔ G₁.Adj = G₂.Adj := by simp [Subgraph.spanningCoe] +lemma mem_of_adj_spanningCoe {v w : V} {s : Set V} (G : SimpleGraph s) + (hadj : G.spanningCoe.Adj v w) : v ∈ s := by aesop + /-- `spanningCoe` is equivalent to `coe` for a subgraph that `IsSpanning`. -/ @[simps] def spanningCoeEquivCoeOfSpanning (G' : Subgraph G) (h : G'.IsSpanning) : diff --git a/Mathlib/Combinatorics/SimpleGraph/Walk.lean b/Mathlib/Combinatorics/SimpleGraph/Walk.lean index ae8224b0ec59d..af48c233e9bcd 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Walk.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Walk.lean @@ -155,6 +155,9 @@ def getVert {u v : V} : G.Walk u v → ℕ → V @[simp] theorem getVert_zero {u v} (w : G.Walk u v) : w.getVert 0 = u := by cases w <;> rfl +@[simp] +theorem getVert_nil (u : V) {i : ℕ} : (@nil _ G u).getVert i = u := rfl + theorem getVert_of_length_le {u v} (w : G.Walk u v) {i : ℕ} (hi : w.length ≤ i) : w.getVert i = v := by induction w generalizing i with @@ -177,11 +180,6 @@ theorem adj_getVert_succ {u v} (w : G.Walk u v) {i : ℕ} (hi : i < w.length) : · simp [getVert, hxy] · exact ih (Nat.succ_lt_succ_iff.1 hi) -lemma getVert_cons_one {u v w} (q : G.Walk v w) (hadj : G.Adj u v) : - (q.cons hadj).getVert 1 = v := by - have : (q.cons hadj).getVert 1 = q.getVert 0 := rfl - simpa [getVert_zero] using this - @[simp] lemma getVert_cons_succ {u v w n} (p : G.Walk v w) (h : G.Adj u v) : (p.cons h).getVert (n + 1) = p.getVert n := rfl @@ -829,10 +827,6 @@ lemma notNilRec_cons {motive : {u w : V} → (p : G.Walk u w) → ¬ p.Nil → S motive (q.cons h) Walk.not_nil_cons) (h' : G.Adj u v) (q' : G.Walk v w) : @Walk.notNilRec _ _ _ _ _ cons _ _ = cons h' q' := by rfl -@[simp] lemma adj_getVert_one {p : G.Walk v w} (hp : ¬ p.Nil) : - G.Adj v (p.getVert 1) := by - simpa using adj_getVert_succ p (by simpa [not_nil_iff_lt_length] using hp : 0 < p.length) - /-- The walk obtained by removing the first `n` darts of a walk. -/ def drop {u v : V} (p : G.Walk u v) (n : ℕ) : G.Walk (p.getVert n) v := match p, n with @@ -840,11 +834,49 @@ def drop {u v : V} (p : G.Walk u v) (n : ℕ) : G.Walk (p.getVert n) v := | p, 0 => p.copy (getVert_zero p).symm rfl | .cons h q, (n + 1) => (q.drop n).copy (getVert_cons_succ _ h).symm rfl -/-- The walk obtained by removing the first dart of a non-nil walk. -/ -def tail (p : G.Walk u v) : G.Walk (p.getVert 1) v := p.drop 1 +/-- The second vertex of a walk, or the only vertex in a nil walk. -/ +abbrev snd (p : G.Walk u v) : V := p.getVert 1 + +@[simp] lemma adj_snd {p : G.Walk v w} (hp : ¬ p.Nil) : + G.Adj v p.snd := by + simpa using adj_getVert_succ p (by simpa [not_nil_iff_lt_length] using hp : 0 < p.length) + +lemma snd_cons {u v w} (q : G.Walk v w) (hadj : G.Adj u v) : + (q.cons hadj).snd = v := by simp + +/-- The penultimate vertex of a walk, or the only vertex in a nil walk. -/ +abbrev penultimate (p : G.Walk u v) : V := p.getVert (p.length - 1) + +@[simp] +lemma penultimate_cons_nil (h : G.Adj u v) : (cons h nil).penultimate = u := rfl + +@[simp] +lemma penultimate_cons_cons {w'} (h : G.Adj u v) (h₂ : G.Adj v w) (p : G.Walk w w') : + (cons h (cons h₂ p)).penultimate = (cons h₂ p).penultimate := rfl + +lemma penultimate_cons_of_not_nil (h : G.Adj u v) (p : G.Walk v w) (hp : ¬ p.Nil) : + (cons h p).penultimate = p.penultimate := + p.notNilRec (by simp) hp h + +@[simp] +lemma penultimate_concat {t u v} (p : G.Walk u v) (h : G.Adj v t) : + (p.concat h).penultimate = v := by simp [penultimate, concat_eq_append, getVert_append] + +@[simp] +lemma adj_penultimate {p : G.Walk v w} (hp : ¬ p.Nil) : + G.Adj p.penultimate w := by + conv => rhs; rw [← getVert_length p] + rw [nil_iff_length_eq] at hp + convert adj_getVert_succ _ _ <;> omega + +/-- The walk obtained by removing the first dart of a walk. A nil walk stays nil. -/ +def tail (p : G.Walk u v) : G.Walk (p.snd) v := p.drop 1 + +@[simp] +lemma tail_nil : (@nil _ G v).tail = .nil := rfl @[simp] -lemma tail_cons_nil (h : G.Adj u v) : (Walk.cons h .nil).tail = .nil := by rfl +lemma tail_cons_nil (h : G.Adj u v) : (Walk.cons h .nil).tail = .nil := rfl lemma tail_cons_eq (h : G.Adj u v) (p : G.Walk v w) : (p.cons h).tail = p.copy (getVert_zero p).symm rfl := by @@ -856,16 +888,26 @@ lemma tail_cons_eq (h : G.Adj u v) (p : G.Walk v w) : @[simps] def firstDart (p : G.Walk v w) (hp : ¬ p.Nil) : G.Dart where fst := v - snd := p.getVert 1 - adj := p.adj_getVert_one hp + snd := p.snd + adj := p.adj_snd hp + +/-- The last dart of a walk. -/ +@[simps] +def lastDart (p : G.Walk v w) (hp : ¬ p.Nil) : G.Dart where + fst := p.penultimate + snd := w + adj := p.adj_penultimate hp lemma edge_firstDart (p : G.Walk v w) (hp : ¬ p.Nil) : - (p.firstDart hp).edge = s(v, p.getVert 1) := rfl + (p.firstDart hp).edge = s(v, p.snd) := rfl + +lemma edge_lastDart (p : G.Walk v w) (hp : ¬ p.Nil) : + (p.lastDart hp).edge = s(p.penultimate, w) := rfl variable {x y : V} -- TODO: rename to u, v, w instead? lemma cons_tail_eq (p : G.Walk x y) (hp : ¬ p.Nil) : - cons (p.adj_getVert_one hp) p.tail = p := by + cons (p.adj_snd hp) p.tail = p := by cases p with | nil => simp only [nil_nil, not_true_eq_false] at hp | cons h q => @@ -1085,7 +1127,7 @@ theorem exists_boundary_dart {u v : V} (p : G.Walk u v) (S : Set V) (uS : u ∈ | .cons hadj q, 0 => simp only [copy_rfl_rfl, getVert_zero] | .cons hadj q, (n + 1) => simp only [copy_cons, getVert_cons_succ]; rfl -@[simp] lemma getVert_tail {u v n} (p : G.Walk u v) (hnp: ¬ p.Nil) : +@[simp] lemma getVert_tail {u v n} (p : G.Walk u v) : p.tail.getVert n = p.getVert (n + 1) := by match p with | .nil => rfl @@ -1121,8 +1163,8 @@ theorem mem_support_iff_exists_getVert {u v w : V} {p : G.Walk v w} : rw [mem_support_iff_exists_getVert] use n - 1 simp only [Nat.sub_le_iff_le_add] - rw [getVert_tail _ hnp, length_tail_add_one hnp] - have : (n - 1 + 1) = n:= by omega + rw [getVert_tail, length_tail_add_one hnp] + have : (n - 1 + 1) = n := by omega rwa [this] termination_by p.length decreasing_by diff --git a/Mathlib/Computability/AkraBazzi/AkraBazzi.lean b/Mathlib/Computability/AkraBazzi/AkraBazzi.lean index 7cd217fadb8cf..0031c9b006e43 100644 --- a/Mathlib/Computability/AkraBazzi/AkraBazzi.lean +++ b/Mathlib/Computability/AkraBazzi/AkraBazzi.lean @@ -843,7 +843,7 @@ lemma isEquivalent_deriv_rpow_p_mul_one_sub_smoothingFn {p : ℝ} (hp : p ≠ 0) (fun z => z ^ (p-1) / (log z ^ 2)) =o[atTop] fun z => z ^ (p-1) / 1 := by simp_rw [div_eq_mul_inv] refine IsBigO.mul_isLittleO (isBigO_refl _ _) - (IsLittleO.inv_rev ?_ (by aesop (add safe Eventually.of_forall))) + (IsLittleO.inv_rev ?_ (by simp)) rw [isLittleO_const_left] refine Or.inr <| Tendsto.comp tendsto_norm_atTop_atTop ?_ exact Tendsto.comp (g := fun z => z ^ 2) @@ -867,7 +867,7 @@ lemma isEquivalent_deriv_rpow_p_mul_one_add_smoothingFn {p : ℝ} (hp : p ≠ 0) (fun z => -(z ^ (p-1) / (log z ^ 2))) =o[atTop] fun z => z ^ (p-1) / 1 := by simp_rw [isLittleO_neg_left, div_eq_mul_inv] refine IsBigO.mul_isLittleO (isBigO_refl _ _) - (IsLittleO.inv_rev ?_ (by aesop (add safe Eventually.of_forall))) + (IsLittleO.inv_rev ?_ (by simp)) rw [isLittleO_const_left] refine Or.inr <| Tendsto.comp tendsto_norm_atTop_atTop ?_ exact Tendsto.comp (g := fun z => z ^ 2) diff --git a/Mathlib/Computability/RegularExpressions.lean b/Mathlib/Computability/RegularExpressions.lean index b816d00833eee..8082aa2d00f52 100644 --- a/Mathlib/Computability/RegularExpressions.lean +++ b/Mathlib/Computability/RegularExpressions.lean @@ -365,7 +365,7 @@ def map (f : α → β) : RegularExpression α → RegularExpression β protected theorem map_pow (f : α → β) (P : RegularExpression α) : ∀ n : ℕ, map f (P ^ n) = map f P ^ n | 0 => by unfold map; rfl - | n + 1 => (congr_arg (· * map f P) (RegularExpression.map_pow f P n) : _) + | n + 1 => (congr_arg (· * map f P) (RegularExpression.map_pow f P n) :) #adaptation_note /-- around nightly-2024-02-25, we need to write `comp x y` in the pattern `comp P Q`, instead of `x * y`. -/ @@ -403,7 +403,7 @@ theorem matches'_map (f : α → β) : exact image_singleton -- Porting note: the following close with last `rw` but not with `simp`? | R + S => by simp only [matches'_map, map, matches'_add]; rw [map_add] - | comp R S => by simp only [matches'_map, map, matches'_mul]; erw [map_mul] + | comp R S => by simp [matches'_map] | star R => by simp_rw [map, matches', matches'_map] rw [Language.kstar_eq_iSup_pow, Language.kstar_eq_iSup_pow] diff --git a/Mathlib/Data/Complex/Abs.lean b/Mathlib/Data/Complex/Abs.lean index 4735e0966f985..7531c38a54ff3 100644 --- a/Mathlib/Data/Complex/Abs.lean +++ b/Mathlib/Data/Complex/Abs.lean @@ -198,9 +198,6 @@ theorem abs_im_div_abs_le_one (z : ℂ) : |z.im / Complex.abs z| ≤ 1 := @[simp, norm_cast] lemma abs_intCast (n : ℤ) : abs n = |↑n| := by rw [← ofReal_intCast, abs_ofReal] -@[deprecated "No deprecation message was provided." (since := "2024-02-14")] -lemma int_cast_abs (n : ℤ) : |↑n| = Complex.abs n := (abs_intCast _).symm - theorem normSq_eq_abs (x : ℂ) : normSq x = (Complex.abs x) ^ 2 := by simp [abs, sq, abs_def, Real.mul_self_sqrt (normSq_nonneg _)] diff --git a/Mathlib/Data/Complex/Module.lean b/Mathlib/Data/Complex/Module.lean index a322de00c4a36..6f3a26e7e4225 100644 --- a/Mathlib/Data/Complex/Module.lean +++ b/Mathlib/Data/Complex/Module.lean @@ -189,13 +189,13 @@ theorem Complex.coe_smul {E : Type*} [AddCommGroup E] [Module ℂ E] (x : ℝ) ( another scalar action of `M` on `E` whenever the action of `ℂ` commutes with the action of `M`. -/ instance (priority := 900) SMulCommClass.complexToReal {M E : Type*} [AddCommGroup E] [Module ℂ E] [SMul M E] [SMulCommClass ℂ M E] : SMulCommClass ℝ M E where - smul_comm r _ _ := (smul_comm (r : ℂ) _ _ : _) + smul_comm r _ _ := smul_comm (r : ℂ) _ _ /-- The scalar action of `ℝ` on a `ℂ`-module `E` induced by `Module.complexToReal` associates with another scalar action of `M` on `E` whenever the action of `ℂ` associates with the action of `M`. -/ instance IsScalarTower.complexToReal {M E : Type*} [AddCommGroup M] [Module ℂ M] [AddCommGroup E] [Module ℂ E] [SMul M E] [IsScalarTower ℂ M E] : IsScalarTower ℝ M E where - smul_assoc r _ _ := (smul_assoc (r : ℂ) _ _ : _) + smul_assoc r _ _ := smul_assoc (r : ℂ) _ _ -- check that the following instance is implied by the one above. example (E : Type*) [AddCommGroup E] [Module ℂ E] : IsScalarTower ℝ ℂ E := inferInstance diff --git a/Mathlib/Data/Countable/Small.lean b/Mathlib/Data/Countable/Small.lean index 72387f0c4a096..b48f61815b730 100644 --- a/Mathlib/Data/Countable/Small.lean +++ b/Mathlib/Data/Countable/Small.lean @@ -18,9 +18,3 @@ universe w v instance (priority := 100) Countable.toSmall (α : Type v) [Countable α] : Small.{w} α := let ⟨_, hf⟩ := exists_injective_nat α small_of_injective hf - -@[deprecated (since := "2024-03-20"), nolint defLemma] -alias small_of_countable := Countable.toSmall - -@[deprecated (since := "2024-03-20"), nolint defLemma] -alias small_of_fintype := Countable.toSmall diff --git a/Mathlib/Data/DFinsupp/Small.lean b/Mathlib/Data/DFinsupp/Small.lean new file mode 100644 index 0000000000000..04278d86c6531 --- /dev/null +++ b/Mathlib/Data/DFinsupp/Small.lean @@ -0,0 +1,28 @@ +/- +Copyright (c) 2025 Sophie. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sophie Morel +-/ +import Mathlib.Data.DFinsupp.Defs +import Mathlib.Logic.Small.Basic + +/-! +# Smallness of the `DFinsupp` type + +Let `π : ι → Type v`. If `ι` and all the `π i` are `w`-small, this provides a `Small.{w}` +instance on `DFinsupp π`. + +-/ + +universe u v w + +variable {ι : Type u} {π : ι → Type v} [∀ i, Zero (π i)] + +section Small + +instance DFinsupp.small [Small.{w} ι] [∀ (i : ι), Small.{w} (π i)] : + Small.{w} (DFinsupp π) := small_of_injective (f := fun x j ↦ x j) (fun f f' eq ↦ by + ext j + exact congr_fun eq j) + +end Small diff --git a/Mathlib/Data/ENNReal/Operations.lean b/Mathlib/Data/ENNReal/Operations.lean index 6cea7623d8216..e97cc16c959c7 100644 --- a/Mathlib/Data/ENNReal/Operations.lean +++ b/Mathlib/Data/ENNReal/Operations.lean @@ -512,13 +512,13 @@ theorem smul_def {M : Type*} [MulAction ℝ≥0∞ M] (c : ℝ≥0) (x : M) : c rfl instance {M N : Type*} [MulAction ℝ≥0∞ M] [MulAction ℝ≥0∞ N] [SMul M N] [IsScalarTower ℝ≥0∞ M N] : - IsScalarTower ℝ≥0 M N where smul_assoc r := (smul_assoc (r : ℝ≥0∞) : _) + IsScalarTower ℝ≥0 M N where smul_assoc r := smul_assoc (r : ℝ≥0∞) instance smulCommClass_left {M N : Type*} [MulAction ℝ≥0∞ N] [SMul M N] [SMulCommClass ℝ≥0∞ M N] : - SMulCommClass ℝ≥0 M N where smul_comm r := (smul_comm (r : ℝ≥0∞) : _) + SMulCommClass ℝ≥0 M N where smul_comm r := smul_comm (r : ℝ≥0∞) instance smulCommClass_right {M N : Type*} [MulAction ℝ≥0∞ N] [SMul M N] [SMulCommClass M ℝ≥0∞ N] : - SMulCommClass M ℝ≥0 N where smul_comm m r := (smul_comm m (r : ℝ≥0∞) : _) + SMulCommClass M ℝ≥0 N where smul_comm m r := smul_comm m (r : ℝ≥0∞) /-- A `DistribMulAction` over `ℝ≥0∞` restricts to a `DistribMulAction` over `ℝ≥0`. -/ noncomputable instance {M : Type*} [AddMonoid M] [DistribMulAction ℝ≥0∞ M] : diff --git a/Mathlib/Data/ENNReal/Real.lean b/Mathlib/Data/ENNReal/Real.lean index 879ef89de544e..bf44809f4e036 100644 --- a/Mathlib/Data/ENNReal/Real.lean +++ b/Mathlib/Data/ENNReal/Real.lean @@ -216,7 +216,7 @@ lemma ofReal_lt_one {p : ℝ} : ENNReal.ofReal p < 1 ↔ p < 1 := by @[simp] lemma ofReal_lt_ofNat {p : ℝ} {n : ℕ} [n.AtLeastTwo] : - ENNReal.ofReal p < no_index (OfNat.ofNat n) ↔ p < OfNat.ofNat n := + ENNReal.ofReal p < ofNat(n) ↔ p < OfNat.ofNat n := ofReal_lt_natCast (NeZero.ne n) @[simp] @@ -232,7 +232,7 @@ lemma one_le_ofReal {p : ℝ} : 1 ≤ ENNReal.ofReal p ↔ 1 ≤ p := by @[simp] lemma ofNat_le_ofReal {n : ℕ} [n.AtLeastTwo] {p : ℝ} : - no_index (OfNat.ofNat n) ≤ ENNReal.ofReal p ↔ OfNat.ofNat n ≤ p := + ofNat(n) ≤ ENNReal.ofReal p ↔ OfNat.ofNat n ≤ p := natCast_le_ofReal (NeZero.ne n) @[simp, norm_cast] @@ -248,7 +248,7 @@ lemma ofReal_le_one {r : ℝ} : ENNReal.ofReal r ≤ 1 ↔ r ≤ 1 := @[simp] lemma ofReal_le_ofNat {r : ℝ} {n : ℕ} [n.AtLeastTwo] : - ENNReal.ofReal r ≤ no_index (OfNat.ofNat n) ↔ r ≤ OfNat.ofNat n := + ENNReal.ofReal r ≤ ofNat(n) ↔ r ≤ OfNat.ofNat n := ofReal_le_natCast @[simp] @@ -263,7 +263,7 @@ lemma one_lt_ofReal {r : ℝ} : 1 < ENNReal.ofReal r ↔ 1 < r := coe_lt_coe.tra @[simp] lemma ofNat_lt_ofReal {n : ℕ} [n.AtLeastTwo] {r : ℝ} : - no_index (OfNat.ofNat n) < ENNReal.ofReal r ↔ OfNat.ofNat n < r := + ofNat(n) < ENNReal.ofReal r ↔ OfNat.ofNat n < r := natCast_lt_ofReal @[simp] @@ -279,7 +279,7 @@ lemma ofReal_eq_one {r : ℝ} : ENNReal.ofReal r = 1 ↔ r = 1 := @[simp] lemma ofReal_eq_ofNat {r : ℝ} {n : ℕ} [n.AtLeastTwo] : - ENNReal.ofReal r = no_index (OfNat.ofNat n) ↔ r = OfNat.ofNat n := + ENNReal.ofReal r = ofNat(n) ↔ r = OfNat.ofNat n := ofReal_eq_natCast (NeZero.ne n) theorem ofReal_sub (p : ℝ) {q : ℝ} (hq : 0 ≤ q) : diff --git a/Mathlib/Data/ENat/Basic.lean b/Mathlib/Data/ENat/Basic.lean index 7f2011f33653d..99492bb51fbaf 100644 --- a/Mathlib/Data/ENat/Basic.lean +++ b/Mathlib/Data/ENat/Basic.lean @@ -104,7 +104,7 @@ def lift (x : ℕ∞) (h : x < ⊤) : ℕ := WithTop.untop x (WithTop.lt_top_iff @[simp] theorem lift_zero : lift 0 (WithTop.coe_lt_top 0) = 0 := rfl @[simp] theorem lift_one : lift 1 (WithTop.coe_lt_top 1) = 1 := rfl @[simp] theorem lift_ofNat (n : ℕ) [n.AtLeastTwo] : - lift (no_index (OfNat.ofNat n)) (WithTop.coe_lt_top n) = OfNat.ofNat n := rfl + lift ofNat(n) (WithTop.coe_lt_top n) = OfNat.ofNat n := rfl @[simp] theorem add_lt_top {a b : ℕ∞} : a + b < ⊤ ↔ a < ⊤ ∧ b < ⊤ := WithTop.add_lt_top @@ -145,9 +145,8 @@ theorem toNat_zero : toNat 0 = 0 := theorem toNat_one : toNat 1 = 1 := rfl --- See note [no_index around OfNat.ofNat] @[simp] -theorem toNat_ofNat (n : ℕ) [n.AtLeastTwo] : toNat (no_index (OfNat.ofNat n)) = n := +theorem toNat_ofNat (n : ℕ) [n.AtLeastTwo] : toNat ofNat(n) = n := rfl @[simp] @@ -164,19 +163,17 @@ theorem recTopCoe_zero {C : ℕ∞ → Sort*} (d : C ⊤) (f : ∀ a : ℕ, C a) theorem recTopCoe_one {C : ℕ∞ → Sort*} (d : C ⊤) (f : ∀ a : ℕ, C a) : @recTopCoe C d f 1 = f 1 := rfl --- See note [no_index around OfNat.ofNat] @[simp] theorem recTopCoe_ofNat {C : ℕ∞ → Sort*} (d : C ⊤) (f : ∀ a : ℕ, C a) (x : ℕ) [x.AtLeastTwo] : - @recTopCoe C d f (no_index (OfNat.ofNat x)) = f (OfNat.ofNat x) := + @recTopCoe C d f ofNat(x) = f (OfNat.ofNat x) := rfl @[simp] theorem top_ne_coe (a : ℕ) : ⊤ ≠ (a : ℕ∞) := nofun --- See note [no_index around OfNat.ofNat] @[simp] -theorem top_ne_ofNat (a : ℕ) [a.AtLeastTwo] : ⊤ ≠ (no_index (OfNat.ofNat a : ℕ∞)) := +theorem top_ne_ofNat (a : ℕ) [a.AtLeastTwo] : ⊤ ≠ (ofNat(a) : ℕ∞) := nofun @[simp] lemma top_ne_zero : (⊤ : ℕ∞) ≠ 0 := nofun @@ -186,9 +183,8 @@ theorem top_ne_ofNat (a : ℕ) [a.AtLeastTwo] : ⊤ ≠ (no_index (OfNat.ofNat a theorem coe_ne_top (a : ℕ) : (a : ℕ∞) ≠ ⊤ := nofun --- See note [no_index around OfNat.ofNat] @[simp] -theorem ofNat_ne_top (a : ℕ) [a.AtLeastTwo] : (no_index (OfNat.ofNat a : ℕ∞)) ≠ ⊤ := +theorem ofNat_ne_top (a : ℕ) [a.AtLeastTwo] : (ofNat(a) : ℕ∞) ≠ ⊤ := nofun @[simp] lemma zero_ne_top : 0 ≠ (⊤ : ℕ∞) := nofun @@ -202,9 +198,8 @@ theorem top_sub_coe (a : ℕ) : (⊤ : ℕ∞) - a = ⊤ := theorem top_sub_one : (⊤ : ℕ∞) - 1 = ⊤ := top_sub_coe 1 --- See note [no_index around OfNat.ofNat] @[simp] -theorem top_sub_ofNat (a : ℕ) [a.AtLeastTwo] : (⊤ : ℕ∞) - (no_index (OfNat.ofNat a)) = ⊤ := +theorem top_sub_ofNat (a : ℕ) [a.AtLeastTwo] : (⊤ : ℕ∞) - ofNat(a) = ⊤ := top_sub_coe a @[simp] @@ -382,7 +377,7 @@ theorem map_zero (f : ℕ → α) : map f 0 = f 0 := rfl theorem map_one (f : ℕ → α) : map f 1 = f 1 := rfl @[simp] -theorem map_ofNat (f : ℕ → α) (n : ℕ) [n.AtLeastTwo] : map f (no_index (OfNat.ofNat n)) = f n := rfl +theorem map_ofNat (f : ℕ → α) (n : ℕ) [n.AtLeastTwo] : map f ofNat(n) = f n := rfl @[simp] lemma map_eq_top_iff {f : ℕ → α} : map f n = ⊤ ↔ n = ⊤ := WithTop.map_eq_top_iff diff --git a/Mathlib/Data/Fin/Basic.lean b/Mathlib/Data/Fin/Basic.lean index 6056e283f70cc..148cf433376f6 100644 --- a/Mathlib/Data/Fin/Basic.lean +++ b/Mathlib/Data/Fin/Basic.lean @@ -79,8 +79,6 @@ def finZeroElim {α : Fin 0 → Sort*} (x : Fin 0) : α x := namespace Fin -@[deprecated (since := "2024-02-15")] alias eq_of_veq := eq_of_val_eq -@[deprecated (since := "2024-02-15")] alias veq_of_eq := val_eq_of_eq @[deprecated (since := "2024-08-13")] alias ne_of_vne := ne_of_val_ne @[deprecated (since := "2024-08-13")] alias vne_of_ne := val_ne_of_ne @@ -151,10 +149,6 @@ section coe theorem val_eq_val (a b : Fin n) : (a : ℕ) = b ↔ a = b := Fin.ext_iff.symm -@[deprecated Fin.ext_iff (since := "2024-02-20")] -theorem eq_iff_veq (a b : Fin n) : a = b ↔ a.1 = b.1 := - Fin.ext_iff - theorem ne_iff_vne (a b : Fin n) : a ≠ b ↔ a.1 ≠ b.1 := Fin.ext_iff.not diff --git a/Mathlib/Data/Finset/Basic.lean b/Mathlib/Data/Finset/Basic.lean index b6901aef2d20c..30dcc372ee2cd 100644 --- a/Mathlib/Data/Finset/Basic.lean +++ b/Mathlib/Data/Finset/Basic.lean @@ -48,7 +48,7 @@ assert_not_exists Multiset.powerset assert_not_exists CompleteLattice -assert_not_exists OrderedCommMonoid +assert_not_exists Monoid open Multiset Subtype Function @@ -532,29 +532,6 @@ variable [DecidableEq α] {s t : Multiset α} theorem toFinset_add (s t : Multiset α) : toFinset (s + t) = toFinset s ∪ toFinset t := Finset.ext <| by simp -@[simp] -theorem toFinset_nsmul (s : Multiset α) : ∀ n ≠ 0, (n • s).toFinset = s.toFinset - | 0, h => by contradiction - | n + 1, _ => by - by_cases h : n = 0 - · rw [h, zero_add, one_nsmul] - · rw [add_nsmul, toFinset_add, one_nsmul, toFinset_nsmul s n h, Finset.union_idempotent] - -theorem toFinset_eq_singleton_iff (s : Multiset α) (a : α) : - s.toFinset = {a} ↔ card s ≠ 0 ∧ s = card s • {a} := by - refine ⟨fun H ↦ ⟨fun h ↦ ?_, ext' fun x ↦ ?_⟩, fun H ↦ ?_⟩ - · rw [card_eq_zero.1 h, toFinset_zero] at H - exact Finset.singleton_ne_empty _ H.symm - · rw [count_nsmul, count_singleton] - by_cases hx : x = a - · simp_rw [hx, ite_true, mul_one, count_eq_card] - intro y hy - rw [← mem_toFinset, H, Finset.mem_singleton] at hy - exact hy.symm - have hx' : x ∉ s := fun h' ↦ hx <| by rwa [← mem_toFinset, H, Finset.mem_singleton] at h' - simp_rw [count_eq_zero_of_not_mem hx', hx, ite_false, Nat.mul_zero] - simpa only [toFinset_nsmul _ _ H.1, toFinset_singleton] using congr($(H.2).toFinset) - @[simp] theorem toFinset_inter (s t : Multiset α) : toFinset (s ∩ t) = toFinset s ∩ toFinset t := Finset.ext <| by simp diff --git a/Mathlib/Data/Finset/Card.lean b/Mathlib/Data/Finset/Card.lean index 648f621a5abac..8c8e3dba2f006 100644 --- a/Mathlib/Data/Finset/Card.lean +++ b/Mathlib/Data/Finset/Card.lean @@ -500,8 +500,6 @@ lemma card_union_eq_card_add_card : #(s ∪ t) = #s + #t ↔ Disjoint s t := by @[simp] alias ⟨_, card_union_of_disjoint⟩ := card_union_eq_card_add_card -@[deprecated (since := "2024-02-09")] alias card_union_eq := card_union_of_disjoint -@[deprecated (since := "2024-02-09")] alias card_disjoint_union := card_union_of_disjoint lemma cast_card_inter [AddGroupWithOne R] : (#(s ∩ t) : R) = #s + #t - #(s ∪ t) := by @@ -612,10 +610,6 @@ theorem card_eq_one : #s = 1 ↔ ∃ a, s = {a} := by cases s simp only [Multiset.card_eq_one, Finset.card, ← val_inj, singleton_val] -theorem _root_.Multiset.toFinset_card_eq_one_iff [DecidableEq α] (s : Multiset α) : - #s.toFinset = 1 ↔ Multiset.card s ≠ 0 ∧ ∃ a : α, s = Multiset.card s • {a} := by - simp_rw [card_eq_one, Multiset.toFinset_eq_singleton_iff, exists_and_left] - theorem exists_eq_insert_iff [DecidableEq α] {s t : Finset α} : (∃ a ∉ s, insert a s = t) ↔ s ⊆ t ∧ #s + 1 = #t := by constructor @@ -676,9 +670,6 @@ theorem one_lt_card_iff_nontrivial : 1 < #s ↔ s.Nontrivial := by rw [← not_iff_not, not_lt, Finset.Nontrivial, ← Set.nontrivial_coe_sort, not_nontrivial_iff_subsingleton, card_le_one_iff_subsingleton_coe, coe_sort_coe] -@[deprecated (since := "2024-02-05")] -alias one_lt_card_iff_nontrivial_coe := one_lt_card_iff_nontrivial - theorem exists_ne_of_one_lt_card (hs : 1 < #s) (a : α) : ∃ b, b ∈ s ∧ b ≠ a := by obtain ⟨x, hx, y, hy, hxy⟩ := Finset.one_lt_card.mp hs by_cases ha : y = a diff --git a/Mathlib/Data/Finset/Defs.lean b/Mathlib/Data/Finset/Defs.lean index ff2e09b64a5eb..e63e94579aaed 100644 --- a/Mathlib/Data/Finset/Defs.lean +++ b/Mathlib/Data/Finset/Defs.lean @@ -63,7 +63,7 @@ assert_not_exists DirectedSystem assert_not_exists CompleteLattice -assert_not_exists OrderedCommMonoid +assert_not_exists Monoid open Multiset Subtype Function @@ -353,7 +353,7 @@ theorem sizeOf_lt_sizeOf_of_mem [SizeOf α] {x : α} {s : Finset α} (hx : x ∈ SizeOf.sizeOf x < SizeOf.sizeOf s := by cases s dsimp [SizeOf.sizeOf, SizeOf.sizeOf, Multiset.sizeOf] - rw [add_comm] + rw [Nat.add_comm] refine lt_trans ?_ (Nat.lt_succ_self _) exact Multiset.sizeOf_lt_sizeOf_of_mem hx diff --git a/Mathlib/Data/Finset/Disjoint.lean b/Mathlib/Data/Finset/Disjoint.lean index fdb75054aba95..4ef7defc5e0a6 100644 --- a/Mathlib/Data/Finset/Disjoint.lean +++ b/Mathlib/Data/Finset/Disjoint.lean @@ -27,7 +27,7 @@ assert_not_exists Multiset.powerset assert_not_exists CompleteLattice -assert_not_exists OrderedCommMonoid +assert_not_exists Monoid open Multiset Subtype Function @@ -134,17 +134,17 @@ theorem coe_disjUnion {s t : Finset α} (h : Disjoint s t) : theorem disjUnion_comm (s t : Finset α) (h : Disjoint s t) : disjUnion s t h = disjUnion t s h.symm := - eq_of_veq <| add_comm _ _ + eq_of_veq <| Multiset.add_comm _ _ @[simp] theorem empty_disjUnion (t : Finset α) (h : Disjoint ∅ t := disjoint_bot_left) : disjUnion ∅ t h = t := - eq_of_veq <| zero_add _ + eq_of_veq <| Multiset.zero_add _ @[simp] theorem disjUnion_empty (s : Finset α) (h : Disjoint s ∅ := disjoint_bot_right) : disjUnion s ∅ h = s := - eq_of_veq <| add_zero _ + eq_of_veq <| Multiset.add_zero _ theorem singleton_disjUnion (a : α) (t : Finset α) (h : Disjoint {a} t) : disjUnion {a} t h = cons a t (disjoint_singleton_left.mp h) := diff --git a/Mathlib/Data/Finset/Empty.lean b/Mathlib/Data/Finset/Empty.lean index 5916c7d3dc88a..60a6034dfb73c 100644 --- a/Mathlib/Data/Finset/Empty.lean +++ b/Mathlib/Data/Finset/Empty.lean @@ -67,7 +67,6 @@ alias ⟨_, Nonempty.coe_sort⟩ := nonempty_coe_sort theorem Nonempty.exists_mem {s : Finset α} (h : s.Nonempty) : ∃ x : α, x ∈ s := h -@[deprecated (since := "2024-03-23")] alias Nonempty.bex := Nonempty.exists_mem theorem Nonempty.mono {s t : Finset α} (hst : s ⊆ t) (hs : s.Nonempty) : t.Nonempty := Set.Nonempty.mono hst hs diff --git a/Mathlib/Data/Finset/Lattice/Basic.lean b/Mathlib/Data/Finset/Lattice/Basic.lean index 14fc97ad84371..79aeaeff8285a 100644 --- a/Mathlib/Data/Finset/Lattice/Basic.lean +++ b/Mathlib/Data/Finset/Lattice/Basic.lean @@ -295,10 +295,6 @@ theorem union_inter_distrib_left (s t u : Finset α) : s ∪ t ∩ u = (s ∪ t) theorem inter_union_distrib_right (s t u : Finset α) : s ∩ t ∪ u = (s ∪ u) ∩ (t ∪ u) := sup_inf_right _ _ _ -@[deprecated (since := "2024-03-22")] alias inter_distrib_left := inter_union_distrib_left -@[deprecated (since := "2024-03-22")] alias inter_distrib_right := union_inter_distrib_right -@[deprecated (since := "2024-03-22")] alias union_distrib_left := union_inter_distrib_left -@[deprecated (since := "2024-03-22")] alias union_distrib_right := inter_union_distrib_right theorem union_union_distrib_left (s t u : Finset α) : s ∪ (t ∪ u) = s ∪ t ∪ (s ∪ u) := sup_sup_distrib_left _ _ _ diff --git a/Mathlib/Data/Finset/Lattice/Lemmas.lean b/Mathlib/Data/Finset/Lattice/Lemmas.lean index c5783884b4b5a..32963dd45c201 100644 --- a/Mathlib/Data/Finset/Lattice/Lemmas.lean +++ b/Mathlib/Data/Finset/Lattice/Lemmas.lean @@ -25,7 +25,7 @@ assert_not_exists Multiset.powerset assert_not_exists CompleteLattice -assert_not_exists OrderedCommMonoid +assert_not_exists Monoid open Multiset Subtype Function diff --git a/Mathlib/Data/Finset/NoncommProd.lean b/Mathlib/Data/Finset/NoncommProd.lean index 712c62a411ec7..ccb90bcd3b3d0 100644 --- a/Mathlib/Data/Finset/NoncommProd.lean +++ b/Mathlib/Data/Finset/NoncommProd.lean @@ -228,6 +228,7 @@ namespace Finset variable [Monoid β] [Monoid γ] +open scoped Function -- required for scoped `on` notation /-- Proof used in definition of `Finset.noncommProd` -/ @[to_additive] diff --git a/Mathlib/Data/Finset/Pairwise.lean b/Mathlib/Data/Finset/Pairwise.lean index 0d7b6a6199507..3e9480a8c16bf 100644 --- a/Mathlib/Data/Finset/Pairwise.lean +++ b/Mathlib/Data/Finset/Pairwise.lean @@ -80,6 +80,8 @@ theorem pairwise_iff_coe_toFinset_pairwise (hn : l.Nodup) (hs : Symmetric r) : letI : IsSymm α r := ⟨hs⟩ rw [coe_toFinset, hn.pairwise_coe] +open scoped Function -- required for scoped `on` notation + theorem pairwise_disjoint_of_coe_toFinset_pairwiseDisjoint {α ι} [SemilatticeInf α] [OrderBot α] [DecidableEq ι] {l : List ι} {f : ι → α} (hl : (l.toFinset : Set ι).PairwiseDisjoint f) (hn : l.Nodup) : l.Pairwise (_root_.Disjoint on f) := diff --git a/Mathlib/Data/Finsupp/Basic.lean b/Mathlib/Data/Finsupp/Basic.lean index 8f9546c197a70..675105181e1fc 100644 --- a/Mathlib/Data/Finsupp/Basic.lean +++ b/Mathlib/Data/Finsupp/Basic.lean @@ -905,17 +905,24 @@ theorem subtypeDomain_apply {a : Subtype p} {v : α →₀ M} : (subtypeDomain p theorem subtypeDomain_zero : subtypeDomain p (0 : α →₀ M) = 0 := rfl -theorem subtypeDomain_eq_zero_iff' {f : α →₀ M} : f.subtypeDomain p = 0 ↔ ∀ x, p x → f x = 0 := by - classical simp_rw [← support_eq_empty, support_subtypeDomain, subtype_eq_empty, - not_mem_support_iff] +theorem subtypeDomain_eq_iff_forall {f g : α →₀ M} : + f.subtypeDomain p = g.subtypeDomain p ↔ ∀ x, p x → f x = g x := by + simp_rw [DFunLike.ext_iff, subtypeDomain_apply, Subtype.forall] + +theorem subtypeDomain_eq_iff {f g : α →₀ M} + (hf : ∀ x ∈ f.support, p x) (hg : ∀ x ∈ g.support, p x) : + f.subtypeDomain p = g.subtypeDomain p ↔ f = g := + subtypeDomain_eq_iff_forall.trans + ⟨fun H ↦ Finsupp.ext fun _a ↦ (em _).elim (H _ <| hf _ ·) fun haf ↦ (em _).elim (H _ <| hg _ ·) + fun hag ↦ (not_mem_support_iff.mp haf).trans (not_mem_support_iff.mp hag).symm, + fun H _ _ ↦ congr($H _)⟩ + +theorem subtypeDomain_eq_zero_iff' {f : α →₀ M} : f.subtypeDomain p = 0 ↔ ∀ x, p x → f x = 0 := + subtypeDomain_eq_iff_forall (g := 0) theorem subtypeDomain_eq_zero_iff {f : α →₀ M} (hf : ∀ x ∈ f.support, p x) : f.subtypeDomain p = 0 ↔ f = 0 := - subtypeDomain_eq_zero_iff'.trans - ⟨fun H => - ext fun x => by - classical exact if hx : p x then H x hx else not_mem_support_iff.1 <| mt (hf x) hx, - fun H x _ => by simp [H]⟩ + subtypeDomain_eq_iff (g := 0) hf (by simp) @[to_additive] theorem prod_subtypeDomain_index [CommMonoid N] {v : α →₀ M} {h : α → M → N} @@ -1504,7 +1511,8 @@ end /-- Given an `AddCommMonoid M` and `s : Set α`, `restrictSupportEquiv s M` is the `Equiv` between the subtype of finitely supported functions with support contained in `s` and the type of finitely supported functions from `s`. -/ -def restrictSupportEquiv (s : Set α) (M : Type*) [AddCommMonoid M] : +-- TODO: add [DecidablePred (· ∈ s)] as an assumption +@[simps] def restrictSupportEquiv (s : Set α) (M : Type*) [AddCommMonoid M] : { f : α →₀ M // ↑f.support ⊆ s } ≃ (s →₀ M) where toFun f := subtypeDomain (· ∈ s) f.1 invFun f := letI := Classical.decPred (· ∈ s); ⟨f.extendDomain, support_extendDomain_subset _⟩ diff --git a/Mathlib/Data/Finsupp/BigOperators.lean b/Mathlib/Data/Finsupp/BigOperators.lean index 170585fa66d4b..ed77d4374edaa 100644 --- a/Mathlib/Data/Finsupp/BigOperators.lean +++ b/Mathlib/Data/Finsupp/BigOperators.lean @@ -70,6 +70,8 @@ theorem Finset.mem_sup_support_iff [Zero M] {s : Finset (ι →₀ M)} {x : ι} x ∈ s.sup Finsupp.support ↔ ∃ f ∈ s, x ∈ f.support := Multiset.mem_sup_map_support_iff +open scoped Function -- required for scoped `on` notation + theorem List.support_sum_eq [AddMonoid M] (l : List (ι →₀ M)) (hl : l.Pairwise (_root_.Disjoint on Finsupp.support)) : l.sum.support = l.foldr (Finsupp.support · ⊔ ·) ∅ := by diff --git a/Mathlib/Data/Finsupp/MonomialOrder/DegLex.lean b/Mathlib/Data/Finsupp/MonomialOrder/DegLex.lean index 45e95c792981c..fc03f56398bf2 100644 --- a/Mathlib/Data/Finsupp/MonomialOrder/DegLex.lean +++ b/Mathlib/Data/Finsupp/MonomialOrder/DegLex.lean @@ -72,6 +72,7 @@ theorem ofDegLex_add [AddCommMonoid α] (a b : DegLex α) : namespace Finsupp +open scoped Function in -- required for scoped `on` notation /-- `Finsupp.DegLex r s` is the homogeneous lexicographic order on `α →₀ M`, where `α` is ordered by `r` and `M` is ordered by `s`. The type synonym `DegLex (α →₀ M)` has an order given by `Finsupp.DegLex (· < ·) (· < ·)`. -/ diff --git a/Mathlib/Data/Finsupp/Multiset.lean b/Mathlib/Data/Finsupp/Multiset.lean index df6435246cdea..0ff33f764dd1d 100644 --- a/Mathlib/Data/Finsupp/Multiset.lean +++ b/Mathlib/Data/Finsupp/Multiset.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.Algebra.Order.Group.Finset import Mathlib.Data.Finsupp.Basic import Mathlib.Data.Finsupp.Order diff --git a/Mathlib/Data/Fintype/Card.lean b/Mathlib/Data/Fintype/Card.lean index 08568f59f0f7d..eaa7a6616ea92 100644 --- a/Mathlib/Data/Fintype/Card.lean +++ b/Mathlib/Data/Fintype/Card.lean @@ -194,7 +194,6 @@ theorem card_eq {α β} [_F : Fintype α] [_G : Fintype β] : card α = card β /-- Note: this lemma is specifically about `Fintype.ofSubsingleton`. For a statement about arbitrary `Fintype` instances, use either `Fintype.card_le_one_iff_subsingleton` or `Fintype.card_unique`. -/ -@[simp] theorem card_ofSubsingleton (a : α) [Subsingleton α] : @Fintype.card _ (ofSubsingleton a) = 1 := rfl @@ -204,7 +203,6 @@ theorem card_unique [Unique α] [h : Fintype α] : Fintype.card α = 1 := /-- Note: this lemma is specifically about `Fintype.ofIsEmpty`. For a statement about arbitrary `Fintype` instances, use `Fintype.card_eq_zero`. -/ -@[simp] theorem card_ofIsEmpty [IsEmpty α] : @Fintype.card α Fintype.ofIsEmpty = 0 := rfl diff --git a/Mathlib/Data/Int/Cast/Basic.lean b/Mathlib/Data/Int/Cast/Basic.lean index ade0aeafd71c2..27cca413964e4 100644 --- a/Mathlib/Data/Int/Cast/Basic.lean +++ b/Mathlib/Data/Int/Cast/Basic.lean @@ -72,7 +72,7 @@ theorem cast_one : ((1 : ℤ) : R) = 1 := by @[simp, norm_cast] theorem cast_neg : ∀ n, ((-n : ℤ) : R) = -n - | (0 : ℕ) => by erw [cast_zero, neg_zero] + | (0 : ℕ) => by simp | (n + 1 : ℕ) => by erw [cast_natCast, cast_negSucc] | -[n+1] => by erw [cast_natCast, cast_negSucc, neg_neg] -- type had `HasLiftT` diff --git a/Mathlib/Data/Int/Cast/Lemmas.lean b/Mathlib/Data/Int/Cast/Lemmas.lean index f414c75984e4a..a9229d8aa827b 100644 --- a/Mathlib/Data/Int/Cast/Lemmas.lean +++ b/Mathlib/Data/Int/Cast/Lemmas.lean @@ -244,10 +244,10 @@ theorem ext_mint {f g : Multiplicative ℤ →* M} (h1 : f (ofAdd 1) = g (ofAdd theorem ext_int {f g : ℤ →* M} (h_neg_one : f (-1) = g (-1)) (h_nat : f.comp Int.ofNatHom.toMonoidHom = g.comp Int.ofNatHom.toMonoidHom) : f = g := by ext (x | x) - · exact (DFunLike.congr_fun h_nat x : _) + · exact (DFunLike.congr_fun h_nat x :) · rw [Int.negSucc_eq, ← neg_one_mul, f.map_mul, g.map_mul] congr 1 - exact mod_cast (DFunLike.congr_fun h_nat (x + 1) : _) + exact mod_cast (DFunLike.congr_fun h_nat (x + 1) :) end MonoidHom @@ -261,7 +261,7 @@ theorem ext_int {f g : ℤ →*₀ M} (h_neg_one : f (-1) = g (-1)) (h_nat : f.comp Int.ofNatHom.toMonoidWithZeroHom = g.comp Int.ofNatHom.toMonoidWithZeroHom) : f = g := toMonoidHom_injective <| MonoidHom.ext_int h_neg_one <| - MonoidHom.ext (DFunLike.congr_fun h_nat : _) + MonoidHom.ext (DFunLike.congr_fun h_nat :) end MonoidWithZeroHom diff --git a/Mathlib/Data/Int/Defs.lean b/Mathlib/Data/Int/Defs.lean index 3addc51ed7381..2b92ec1456668 100644 --- a/Mathlib/Data/Int/Defs.lean +++ b/Mathlib/Data/Int/Defs.lean @@ -91,9 +91,6 @@ protected theorem ofNat_add_one_out (n : ℕ) : ↑n + (1 : ℤ) = ↑(succ n) : @[simp] lemma ofNat_eq_natCast (n : ℕ) : Int.ofNat n = n := rfl -@[deprecated ofNat_eq_natCast (since := "2024-03-24")] -protected lemma natCast_eq_ofNat (n : ℕ) : ↑n = Int.ofNat n := rfl - @[norm_cast] lemma natCast_inj {m n : ℕ} : (m : ℤ) = (n : ℤ) ↔ m = n := ofNat_inj @[simp, norm_cast] lemma natAbs_cast (n : ℕ) : natAbs ↑n = n := rfl @@ -222,7 +219,7 @@ than `b`, and the `pred` of a number less than `b`. -/ where /-- The positive case of `Int.inductionOn'`. -/ pos : ∀ n : ℕ, C (b + n) - | 0 => cast (by erw [Int.add_zero]) H0 + | 0 => cast (by simp) H0 | n+1 => cast (by rw [Int.add_assoc]; rfl) <| Hs _ (Int.le_add_of_nonneg_right (ofNat_nonneg _)) (pos n) diff --git a/Mathlib/Data/List/Basic.lean b/Mathlib/Data/List/Basic.lean index 4c99fc84425e5..6b908db84ff95 100644 --- a/Mathlib/Data/List/Basic.lean +++ b/Mathlib/Data/List/Basic.lean @@ -73,7 +73,6 @@ theorem _root_.Decidable.List.eq_or_ne_mem_of_mem [DecidableEq α] lemma mem_pair {a b c : α} : a ∈ [b, c] ↔ a = b ∨ a = c := by rw [mem_cons, mem_singleton] -@[deprecated (since := "2024-03-23")] alias mem_split := append_of_mem -- The simpNF linter says that the LHS can be simplified via `List.mem_map`. -- However this is a higher priority lemma. @@ -329,7 +328,7 @@ theorem getLast_replicate_succ (m : ℕ) (a : α) : lemma getLast_filter' {p : α → Bool} : ∀ (l : List α) (hlp : l.filter p ≠ []), p (l.getLast (hlp <| ·.symm ▸ rfl)) = true → (l.filter p).getLast hlp = l.getLast (hlp <| ·.symm ▸ rfl) - | [a], h, h' => by rw [List.getLast_singleton'] at h'; simp [List.filter_cons, h'] + | [a], h, h' => by simp | a :: b :: as, h, h' => by rw [List.getLast_cons_cons] at h' ⊢ simp only [List.filter_cons (x := a)] at h ⊢ @@ -872,14 +871,9 @@ theorem flatMap_pure_eq_map (f : α → β) (l : List α) : l.flatMap (pure ∘ @[deprecated (since := "2024-10-16")] alias bind_pure_eq_map := flatMap_pure_eq_map -set_option linter.deprecated false in -@[deprecated flatMap_pure_eq_map (since := "2024-03-24")] -theorem bind_ret_eq_map (f : α → β) (l : List α) : l.bind (List.ret ∘ f) = map f l := - bind_pure_eq_map f l - theorem flatMap_congr {l : List α} {f g : α → List β} (h : ∀ x ∈ l, f x = g x) : List.flatMap l f = List.flatMap l g := - (congr_arg List.flatten <| map_congr_left h : _) + (congr_arg List.flatten <| map_congr_left h :) @[deprecated (since := "2024-10-16")] alias bind_congr := flatMap_congr diff --git a/Mathlib/Data/List/Defs.lean b/Mathlib/Data/List/Defs.lean index b340e9b270a42..9718d8cc7eb2c 100644 --- a/Mathlib/Data/List/Defs.lean +++ b/Mathlib/Data/List/Defs.lean @@ -504,8 +504,6 @@ alias ⟨eq_or_mem_of_mem_cons, _⟩ := mem_cons theorem not_exists_mem_nil (p : α → Prop) : ¬∃ x ∈ @nil α, p x := fun ⟨_, hx, _⟩ => List.not_mem_nil _ hx -@[deprecated (since := "2024-03-23")] alias not_bex_nil := not_exists_mem_nil -@[deprecated (since := "2024-03-23")] alias bex_cons := exists_mem_cons @[deprecated (since := "2024-08-10")] alias length_le_of_sublist := Sublist.length_le diff --git a/Mathlib/Data/List/Infix.lean b/Mathlib/Data/List/Infix.lean index ddab47721a8c6..942f206a4afa8 100644 --- a/Mathlib/Data/List/Infix.lean +++ b/Mathlib/Data/List/Infix.lean @@ -106,7 +106,6 @@ instance decidableInfix [DecidableEq α] : ∀ l₁ l₂ : List α, Decidable (l theorem cons_prefix_iff : a :: l₁ <+: b :: l₂ ↔ a = b ∧ l₁ <+: l₂ := by simp -@[deprecated (since := "2024-03-26")] alias IsPrefix.filter_map := IsPrefix.filterMap protected theorem IsPrefix.reduceOption {l₁ l₂ : List (Option α)} (h : l₁ <+: l₂) : l₁.reduceOption <+: l₂.reduceOption := diff --git a/Mathlib/Data/List/Pairwise.lean b/Mathlib/Data/List/Pairwise.lean index be4f87bb84be6..be0725e086b7b 100644 --- a/Mathlib/Data/List/Pairwise.lean +++ b/Mathlib/Data/List/Pairwise.lean @@ -49,15 +49,6 @@ theorem Pairwise.forall (hR : Symmetric R) (hl : l.Pairwise R) : theorem Pairwise.set_pairwise (hl : Pairwise R l) (hr : Symmetric R) : { x | x ∈ l }.Pairwise R := hl.forall hr --- Porting note: Duplicate of `pairwise_map` but with `f` explicit. -@[deprecated "No deprecation message was provided." (since := "2024-02-25")] -theorem pairwise_map' (f : β → α) : - ∀ {l : List β}, Pairwise R (map f l) ↔ Pairwise (R on f) l - | [] => by simp only [map, Pairwise.nil] - | b :: l => by - simp only [map, pairwise_cons, mem_map, forall_exists_index, and_imp, - forall_apply_eq_imp_iff₂, pairwise_map] - theorem pairwise_of_reflexive_of_forall_ne {l : List α} {r : α → α → Prop} (hr : Reflexive r) (h : ∀ a ∈ l, ∀ b ∈ l, a ≠ b → r a b) : l.Pairwise r := by rw [pairwise_iff_forall_sublist] diff --git a/Mathlib/Data/List/Sublists.lean b/Mathlib/Data/List/Sublists.lean index aca81a62467f5..cf9436cf68118 100644 --- a/Mathlib/Data/List/Sublists.lean +++ b/Mathlib/Data/List/Sublists.lean @@ -184,11 +184,6 @@ theorem map_pure_sublist_sublists (l : List α) : map pure l <+ sublists l := by singleton_sublist.2 <| mem_map.2 ⟨[], mem_sublists.2 (nil_sublist _), by rfl⟩).trans ((append_sublist_append_right _).2 ih) -set_option linter.deprecated false in -@[deprecated map_pure_sublist_sublists (since := "2024-03-24")] -theorem map_ret_sublist_sublists (l : List α) : map List.ret l <+ sublists l := - map_pure_sublist_sublists l - /-! ### sublistsLen -/ /-- Auxiliary function to construct the list of all sublists of a given length. Given an diff --git a/Mathlib/Data/MLList/BestFirst.lean b/Mathlib/Data/MLList/BestFirst.lean index 430ac91fa3323..4f0ae77aca908 100644 --- a/Mathlib/Data/MLList/BestFirst.lean +++ b/Mathlib/Data/MLList/BestFirst.lean @@ -307,7 +307,7 @@ variable {m : Type → Type} {α : Type} [Monad m] [Alternative m] [LinearOrder for simple use cases. -/ local instance instOrderBotEq {a : α} : OrderBot { x : α // x = a } where bot := ⟨a, rfl⟩ - bot_le := by aesop + bot_le := by simp /-- A lazy list recording the best first search of a graph generated by a function diff --git a/Mathlib/Data/Matrix/ColumnRowPartitioned.lean b/Mathlib/Data/Matrix/ColumnRowPartitioned.lean index c583ebb1d3b5e..d4d5448a54f8f 100644 --- a/Mathlib/Data/Matrix/ColumnRowPartitioned.lean +++ b/Mathlib/Data/Matrix/ColumnRowPartitioned.lean @@ -125,8 +125,7 @@ lemma fromRows_toRows (A : Matrix (m₁ ⊕ m₂) n R) : fromRows A.toRows₁ A. lemma fromRows_inj : Function.Injective2 (@fromRows R m₁ m₂ n) := by intros x1 x2 y1 y2 - simp only [funext_iff, ← Matrix.ext_iff] - aesop + simp [← Matrix.ext_iff] lemma fromCols_inj : Function.Injective2 (@fromCols R m n₁ n₂) := by intros x1 x2 y1 y2 diff --git a/Mathlib/Data/Matrix/ConjTranspose.lean b/Mathlib/Data/Matrix/ConjTranspose.lean index 5f5843664ca6b..4b856340799ca 100644 --- a/Mathlib/Data/Matrix/ConjTranspose.lean +++ b/Mathlib/Data/Matrix/ConjTranspose.lean @@ -163,17 +163,15 @@ theorem conjTranspose_eq_natCast [DecidableEq n] [Semiring α] [StarRing α] (Function.Involutive.eq_iff conjTranspose_conjTranspose).trans <| by rw [conjTranspose_natCast] --- See note [no_index around OfNat.ofNat] @[simp] theorem conjTranspose_ofNat [DecidableEq n] [Semiring α] [StarRing α] (d : ℕ) [d.AtLeastTwo] : - (no_index (OfNat.ofNat d) : Matrix n n α)ᴴ = OfNat.ofNat d := + (ofNat(d) : Matrix n n α)ᴴ = OfNat.ofNat d := conjTranspose_natCast _ --- See note [no_index around OfNat.ofNat] @[simp] theorem conjTranspose_eq_ofNat [DecidableEq n] [Semiring α] [StarRing α] {M : Matrix n n α} {d : ℕ} [d.AtLeastTwo] : - Mᴴ = no_index (OfNat.ofNat d) ↔ M = OfNat.ofNat d := + Mᴴ = ofNat(d) ↔ M = OfNat.ofNat d := conjTranspose_eq_natCast @[simp] @@ -239,11 +237,10 @@ theorem conjTranspose_natCast_smul [Semiring R] [AddCommMonoid α] [StarAddMonoi (c : ℕ) (M : Matrix m n α) : ((c : R) • M)ᴴ = (c : R) • Mᴴ := Matrix.ext <| by simp --- See note [no_index around OfNat.ofNat] @[simp] theorem conjTranspose_ofNat_smul [Semiring R] [AddCommMonoid α] [StarAddMonoid α] [Module R α] (c : ℕ) [c.AtLeastTwo] (M : Matrix m n α) : - ((no_index (OfNat.ofNat c : R)) • M)ᴴ = (OfNat.ofNat c : R) • Mᴴ := + ((ofNat(c) : R) • M)ᴴ = (OfNat.ofNat c : R) • Mᴴ := conjTranspose_natCast_smul c M @[simp] @@ -256,11 +253,10 @@ theorem conjTranspose_inv_natCast_smul [DivisionSemiring R] [AddCommMonoid α] [ [Module R α] (c : ℕ) (M : Matrix m n α) : ((c : R)⁻¹ • M)ᴴ = (c : R)⁻¹ • Mᴴ := Matrix.ext <| by simp --- See note [no_index around OfNat.ofNat] @[simp] theorem conjTranspose_inv_ofNat_smul [DivisionSemiring R] [AddCommMonoid α] [StarAddMonoid α] [Module R α] (c : ℕ) [c.AtLeastTwo] (M : Matrix m n α) : - ((no_index (OfNat.ofNat c : R))⁻¹ • M)ᴴ = (OfNat.ofNat c : R)⁻¹ • Mᴴ := + ((ofNat(c) : R)⁻¹ • M)ᴴ = (OfNat.ofNat c : R)⁻¹ • Mᴴ := conjTranspose_inv_natCast_smul c M @[simp] diff --git a/Mathlib/Data/Matrix/Defs.lean b/Mathlib/Data/Matrix/Defs.lean index d85e8a18a4b44..cfe899da9503e 100644 --- a/Mathlib/Data/Matrix/Defs.lean +++ b/Mathlib/Data/Matrix/Defs.lean @@ -465,7 +465,7 @@ theorem reindex_symm (eₘ : m ≃ l) (eₙ : n ≃ o) : theorem reindex_trans {l₂ o₂ : Type*} (eₘ : m ≃ l) (eₙ : n ≃ o) (eₘ₂ : l ≃ l₂) (eₙ₂ : o ≃ o₂) : (reindex eₘ eₙ).trans (reindex eₘ₂ eₙ₂) = (reindex (eₘ.trans eₘ₂) (eₙ.trans eₙ₂) : Matrix m n α ≃ _) := - Equiv.ext fun A => (A.submatrix_submatrix eₘ.symm eₙ.symm eₘ₂.symm eₙ₂.symm : _) + Equiv.ext fun A => (A.submatrix_submatrix eₘ.symm eₙ.symm eₘ₂.symm eₙ₂.symm :) theorem transpose_reindex (eₘ : m ≃ l) (eₙ : n ≃ o) (M : Matrix m n α) : (reindex eₘ eₙ M)ᵀ = reindex eₙ eₘ Mᵀ := diff --git a/Mathlib/Data/Matrix/Diagonal.lean b/Mathlib/Data/Matrix/Diagonal.lean index a1ae509a8b378..89efdc3f4b838 100644 --- a/Mathlib/Data/Matrix/Diagonal.lean +++ b/Mathlib/Data/Matrix/Diagonal.lean @@ -117,13 +117,11 @@ theorem diagonal_natCast [Zero α] [NatCast α] (m : ℕ) : diagonal (fun _ : n @[norm_cast] theorem diagonal_natCast' [Zero α] [NatCast α] (m : ℕ) : diagonal ((m : n → α)) = m := rfl --- See note [no_index around OfNat.ofNat] theorem diagonal_ofNat [Zero α] [NatCast α] (m : ℕ) [m.AtLeastTwo] : - diagonal (fun _ : n => no_index (OfNat.ofNat m : α)) = OfNat.ofNat m := rfl + diagonal (fun _ : n => (ofNat(m) : α)) = OfNat.ofNat m := rfl --- See note [no_index around OfNat.ofNat] theorem diagonal_ofNat' [Zero α] [NatCast α] (m : ℕ) [m.AtLeastTwo] : - diagonal (no_index (OfNat.ofNat m : n → α)) = OfNat.ofNat m := rfl + diagonal (ofNat(m) : n → α) = OfNat.ofNat m := rfl instance [Zero α] [IntCast α] : IntCast (Matrix n n α) where intCast m := diagonal fun _ => m @@ -146,10 +144,9 @@ protected theorem map_natCast [AddMonoidWithOne α] [AddMonoidWithOne β] (d : Matrix n n α).map f = diagonal (fun _ => f d) := diagonal_map h --- See note [no_index around OfNat.ofNat] protected theorem map_ofNat [AddMonoidWithOne α] [AddMonoidWithOne β] {f : α → β} (h : f 0 = 0) (d : ℕ) [d.AtLeastTwo] : - (no_index (OfNat.ofNat d) : Matrix n n α).map f = diagonal (fun _ => f (OfNat.ofNat d)) := + (ofNat(d) : Matrix n n α).map f = diagonal (fun _ => f (OfNat.ofNat d)) := diagonal_map h protected theorem map_intCast [AddGroupWithOne α] [AddGroupWithOne β] @@ -310,17 +307,15 @@ theorem transpose_eq_natCast [DecidableEq n] [AddMonoidWithOne α] {M : Matrix n Mᵀ = d ↔ M = d := transpose_eq_diagonal --- See note [no_index around OfNat.ofNat] @[simp] theorem transpose_ofNat [DecidableEq n] [AddMonoidWithOne α] (d : ℕ) [d.AtLeastTwo] : - (no_index (OfNat.ofNat d) : Matrix n n α)ᵀ = OfNat.ofNat d := + (ofNat(d) : Matrix n n α)ᵀ = OfNat.ofNat d := transpose_natCast _ --- See note [no_index around OfNat.ofNat] @[simp] theorem transpose_eq_ofNat [DecidableEq n] [AddMonoidWithOne α] {M : Matrix n n α} {d : ℕ} [d.AtLeastTwo] : - Mᵀ = no_index (OfNat.ofNat d) ↔ M = OfNat.ofNat d := + Mᵀ = ofNat(d) ↔ M = OfNat.ofNat d := transpose_eq_diagonal @[simp] diff --git a/Mathlib/Data/Matrix/Kronecker.lean b/Mathlib/Data/Matrix/Kronecker.lean index 1d108492e89cf..0d5b9fc77bc99 100644 --- a/Mathlib/Data/Matrix/Kronecker.lean +++ b/Mathlib/Data/Matrix/Kronecker.lean @@ -313,12 +313,12 @@ theorem natCast_kronecker [NonAssocSemiring α] [DecidableEq l] (a : ℕ) (B : M simp [(Nat.cast_commute a _).eq] theorem kronecker_ofNat [Semiring α] [DecidableEq n] (A : Matrix l m α) (b : ℕ) [b.AtLeastTwo] : - A ⊗ₖ (no_index (OfNat.ofNat b) : Matrix n n α) = + A ⊗ₖ (ofNat(b) : Matrix n n α) = blockDiagonal fun _ => A <• (OfNat.ofNat b : α) := kronecker_diagonal _ _ theorem ofNat_kronecker [Semiring α] [DecidableEq l] (a : ℕ) [a.AtLeastTwo] (B : Matrix m n α) : - (no_index (OfNat.ofNat a) : Matrix l l α) ⊗ₖ B = + (ofNat(a) : Matrix l l α) ⊗ₖ B = Matrix.reindex (.prodComm _ _) (.prodComm _ _) (blockDiagonal fun _ => (OfNat.ofNat a : α) • B) := diagonal_kronecker _ _ diff --git a/Mathlib/Data/Matrix/Mul.lean b/Mathlib/Data/Matrix/Mul.lean index 96502a2383140..3d4f563be4268 100644 --- a/Mathlib/Data/Matrix/Mul.lean +++ b/Mathlib/Data/Matrix/Mul.lean @@ -814,16 +814,14 @@ theorem vecMul_natCast (x : ℕ) (v : m → α) : v ᵥ* x = MulOpposite.op (x : vecMul_diagonal_const _ _ --- See note [no_index around OfNat.ofNat] @[simp] theorem ofNat_mulVec (x : ℕ) [x.AtLeastTwo] (v : m → α) : - OfNat.ofNat (no_index x) *ᵥ v = (OfNat.ofNat x : α) • v := + ofNat(x) *ᵥ v = (OfNat.ofNat x : α) • v := natCast_mulVec _ _ --- See note [no_index around OfNat.ofNat] @[simp] theorem vecMul_ofNat (x : ℕ) [x.AtLeastTwo] (v : m → α) : - v ᵥ* OfNat.ofNat (no_index x) = MulOpposite.op (OfNat.ofNat x : α) • v := + v ᵥ* ofNat(x) = MulOpposite.op (OfNat.ofNat x : α) • v := vecMul_natCast _ _ end NonAssocSemiring diff --git a/Mathlib/Data/Matrix/Notation.lean b/Mathlib/Data/Matrix/Notation.lean index 66a14eeaec9ed..e29b479e13304 100644 --- a/Mathlib/Data/Matrix/Notation.lean +++ b/Mathlib/Data/Matrix/Notation.lean @@ -377,7 +377,7 @@ theorem submatrix_cons_row (A : Matrix m' n' α) (i : m') (row : Fin m → m') ( @[simp] theorem submatrix_updateRow_succAbove (A : Matrix (Fin m.succ) n' α) (v : n' → α) (f : o' → n') (i : Fin m.succ) : (A.updateRow i v).submatrix i.succAbove f = A.submatrix i.succAbove f := - ext fun r s => (congr_fun (updateRow_ne (Fin.succAbove_ne i r) : _ = A _) (f s) : _) + ext fun r s => (congr_fun (updateRow_ne (Fin.succAbove_ne i r) : _ = A _) (f s) :) /-- Updating a column then removing it is the same as removing it. -/ @[simp] @@ -418,15 +418,13 @@ theorem natCast_fin_three (n : ℕ) : ext i j fin_cases i <;> fin_cases j <;> rfl --- See note [no_index around OfNat.ofNat] theorem ofNat_fin_two (n : ℕ) [n.AtLeastTwo] : - (no_index (OfNat.ofNat n) : Matrix (Fin 2) (Fin 2) α) = + (ofNat(n) : Matrix (Fin 2) (Fin 2) α) = !![OfNat.ofNat n, 0; 0, OfNat.ofNat n] := natCast_fin_two _ --- See note [no_index around OfNat.ofNat] theorem ofNat_fin_three (n : ℕ) [n.AtLeastTwo] : - (no_index (OfNat.ofNat n) : Matrix (Fin 3) (Fin 3) α) = + (ofNat(n) : Matrix (Fin 3) (Fin 3) α) = !![OfNat.ofNat n, 0, 0; 0, OfNat.ofNat n, 0; 0, 0, OfNat.ofNat n] := natCast_fin_three _ diff --git a/Mathlib/Data/Matroid/Closure.lean b/Mathlib/Data/Matroid/Closure.lean index cb14ccdda92a3..534a3d8ddbc12 100644 --- a/Mathlib/Data/Matroid/Closure.lean +++ b/Mathlib/Data/Matroid/Closure.lean @@ -154,7 +154,7 @@ lemma closure_subset_ground (M : Matroid α) (X : Set α) : M.closure X ⊆ M.E simp_rw [closure_def, inter_assoc, inter_self] lemma inter_ground_subset_closure (M : Matroid α) (X : Set α) : X ∩ M.E ⊆ M.closure X := by - simp_rw [closure_def, subset_sInter_iff]; aesop + simp_rw [closure_def, subset_sInter_iff]; simp lemma mem_closure_iff_forall_mem_flat (X : Set α) (hX : X ⊆ M.E := by aesop_mat) : e ∈ M.closure X ↔ ∀ F, M.Flat F → X ⊆ F → e ∈ F := by diff --git a/Mathlib/Data/Matroid/Sum.lean b/Mathlib/Data/Matroid/Sum.lean index f3e98ee11cd2c..99c771b42c302 100644 --- a/Mathlib/Data/Matroid/Sum.lean +++ b/Mathlib/Data/Matroid/Sum.lean @@ -185,6 +185,8 @@ end sum' section disjointSigma +open scoped Function -- required for scoped `on` notation + variable {α ι : Type*} {M : ι → Matroid α} /-- The sum of an indexed collection of matroids on `α` with pairwise disjoint ground sets, diff --git a/Mathlib/Data/Multiset/Basic.lean b/Mathlib/Data/Multiset/Basic.lean index f7425d084e924..bb263451f18a0 100644 --- a/Mathlib/Data/Multiset/Basic.lean +++ b/Mathlib/Data/Multiset/Basic.lean @@ -3,14 +3,13 @@ Copyright (c) 2015 Microsoft Corporation. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ -import Mathlib.Algebra.Group.Hom.Defs -import Mathlib.Algebra.Group.Nat.Basic -import Mathlib.Algebra.Order.Sub.Unbundled.Basic +import Mathlib.Data.List.Count import Mathlib.Data.List.Perm.Basic import Mathlib.Data.List.Perm.Lattice import Mathlib.Data.List.Perm.Subperm import Mathlib.Data.Set.List import Mathlib.Order.Hom.Basic +import Mathlib.Order.MinMax /-! # Multisets @@ -33,8 +32,8 @@ lists by permutation. This gives them computational content. occurrences of `a` in `s` and `t`. -/ --- No bundled ordered algebra should be required -assert_not_exists OrderedCommMonoid +-- No algebra should be required +assert_not_exists Monoid universe v @@ -586,6 +585,8 @@ end /-! ### Additive monoid -/ +section add +variable {s t u : Multiset α} /-- The sum of two multisets is the lift of the list append operation. This adds the multiplicities of each element, @@ -605,28 +606,33 @@ theorem coe_add (s t : List α) : (s + t : Multiset α) = (s ++ t : List α) := theorem singleton_add (a : α) (s : Multiset α) : {a} + s = a ::ₘ s := rfl -private theorem add_le_add_iff_left' {s t u : Multiset α} : s + t ≤ s + u ↔ t ≤ u := +protected lemma add_le_add_iff_left : s + t ≤ s + u ↔ t ≤ u := Quotient.inductionOn₃ s t u fun _ _ _ => subperm_append_left _ -instance : AddLeftMono (Multiset α) := - ⟨fun _s _t _u => add_le_add_iff_left'.2⟩ +protected lemma add_le_add_iff_right : s + u ≤ t + u ↔ s ≤ t := + Quotient.inductionOn₃ s t u fun _ _ _ => subperm_append_right _ -instance : AddLeftReflectLE (Multiset α) := - ⟨fun _s _t _u => add_le_add_iff_left'.1⟩ +protected alias ⟨le_of_add_le_add_left, add_le_add_left⟩ := Multiset.add_le_add_iff_left +protected alias ⟨le_of_add_le_add_right, add_le_add_right⟩ := Multiset.add_le_add_iff_right -instance instAddCommMonoid : AddCancelCommMonoid (Multiset α) where - add_comm := fun s t => Quotient.inductionOn₂ s t fun _ _ => Quot.sound perm_append_comm - add_assoc := fun s₁ s₂ s₃ => - Quotient.inductionOn₃ s₁ s₂ s₃ fun l₁ l₂ l₃ => congr_arg _ <| append_assoc l₁ l₂ l₃ - zero_add := fun s => Quot.inductionOn s fun _ => rfl - add_zero := fun s => Quotient.inductionOn s fun l => congr_arg _ <| append_nil l - add_left_cancel := fun _ _ _ h => - le_antisymm (Multiset.add_le_add_iff_left'.mp h.le) (Multiset.add_le_add_iff_left'.mp h.ge) - nsmul := nsmulRec +protected lemma add_comm (s t : Multiset α) : s + t = t + s := + Quotient.inductionOn₂ s t fun _ _ ↦ Quot.sound perm_append_comm -theorem le_add_right (s t : Multiset α) : s ≤ s + t := by simpa using add_le_add_left (zero_le t) s +protected lemma add_assoc (s t u : Multiset α) : s + t + u = s + (t + u) := + Quotient.inductionOn₃ s t u fun _ _ _ ↦ congr_arg _ <| append_assoc .. -theorem le_add_left (s t : Multiset α) : s ≤ t + s := by simpa using add_le_add_right (zero_le t) s +@[simp, nolint simpNF] -- We want to use this lemma earlier than `zero_add` +protected lemma zero_add (s : Multiset α) : 0 + s = s := Quotient.inductionOn s fun _ ↦ rfl + +@[simp, nolint simpNF] -- We want to use this lemma earlier than `add_zero` +protected lemma add_zero (s : Multiset α) : s + 0 = s := + Quotient.inductionOn s fun l ↦ congr_arg _ <| append_nil l + +lemma le_add_right (s t : Multiset α) : s ≤ s + t := by + simpa using Multiset.add_le_add_left (zero_le t) + +lemma le_add_left (s t : Multiset α) : s ≤ t + s := by + simpa using Multiset.add_le_add_right (zero_le t) lemma subset_add_left {s t : Multiset α} : s ⊆ s + t := subset_of_le <| le_add_right s t @@ -641,45 +647,24 @@ theorem le_iff_exists_add {s t : Multiset α} : s ≤ t ↔ ∃ u, t = s + u := @[simp] theorem cons_add (a : α) (s t : Multiset α) : a ::ₘ s + t = a ::ₘ (s + t) := by - rw [← singleton_add, ← singleton_add, add_assoc] + rw [← singleton_add, ← singleton_add, Multiset.add_assoc] @[simp] theorem add_cons (a : α) (s t : Multiset α) : s + a ::ₘ t = a ::ₘ (s + t) := by - rw [add_comm, cons_add, add_comm] + rw [Multiset.add_comm, cons_add, Multiset.add_comm] @[simp] theorem mem_add {a : α} {s t : Multiset α} : a ∈ s + t ↔ a ∈ s ∨ a ∈ t := Quotient.inductionOn₂ s t fun _l₁ _l₂ => mem_append -theorem mem_of_mem_nsmul {a : α} {s : Multiset α} {n : ℕ} (h : a ∈ n • s) : a ∈ s := by - induction' n with n ih - · rw [zero_nsmul] at h - exact absurd h (not_mem_zero _) - · rw [succ_nsmul, mem_add] at h - exact h.elim ih id - -@[simp] -theorem mem_nsmul {a : α} {s : Multiset α} {n : ℕ} : a ∈ n • s ↔ n ≠ 0 ∧ a ∈ s := by - refine ⟨fun ha ↦ ⟨?_, mem_of_mem_nsmul ha⟩, fun h => ?_⟩ - · rintro rfl - simp [zero_nsmul] at ha - obtain ⟨n, rfl⟩ := exists_eq_succ_of_ne_zero h.1 - rw [succ_nsmul, mem_add] - exact Or.inr h.2 - -lemma mem_nsmul_of_ne_zero {a : α} {s : Multiset α} {n : ℕ} (h0 : n ≠ 0) : a ∈ n • s ↔ a ∈ s := by - simp [*] - -theorem nsmul_cons {s : Multiset α} (n : ℕ) (a : α) : - n • (a ::ₘ s) = n • ({a} : Multiset α) + n • s := by - rw [← singleton_add, nsmul_add] +end add /-! ### Cardinality -/ /-- The cardinality of a multiset is the sum of the multiplicities of all its elements, or simply the length of the underlying list. -/ -def card (s : Multiset α) : ℕ := Quot.liftOn s length fun _l₁ _l₂ => Perm.length_eq +def card : Multiset α → ℕ := Quot.lift length fun _l₁ _l₂ => Perm.length_eq @[simp] theorem coe_card (l : List α) : card (l : Multiset α) = length l := @@ -693,18 +678,9 @@ theorem length_toList (s : Multiset α) : s.toList.length = card s := by theorem card_zero : @card α 0 = 0 := rfl -@[simp] lemma card_add (s t : Multiset α) : card (s + t) = card s + card t := - Quotient.inductionOn₂ s t length_append - -/-- `Multiset.card` bundled as a monoid hom. -/ -@[simps] -def cardHom : Multiset α →+ ℕ where - toFun := card - map_zero' := card_zero - map_add' := card_add - @[simp] -lemma card_nsmul (s : Multiset α) (n : ℕ) : card (n • s) = n * card s := cardHom.map_nsmul .. +theorem card_add (s t : Multiset α) : card (s + t) = card s + card t := + Quotient.inductionOn₂ s t length_append @[simp] theorem card_cons (a : α) (s : Multiset α) : card (a ::ₘ s) = card s + 1 := @@ -712,7 +688,7 @@ theorem card_cons (a : α) (s : Multiset α) : card (a ::ₘ s) = card s + 1 := @[simp] theorem card_singleton (a : α) : card ({a} : Multiset α) = 1 := by - simp only [← cons_zero, card_zero, eq_self_iff_true, zero_add, card_cons] + simp only [← cons_zero, card_zero, eq_self_iff_true, Multiset.zero_add, card_cons] theorem card_pair (a b : α) : card {a, b} = 2 := by rw [insert_eq_cons, card_cons, card_singleton] @@ -836,13 +812,6 @@ theorem coe_replicate (n : ℕ) (a : α) : (List.replicate n a : Multiset α) = theorem replicate_add (m n : ℕ) (a : α) : replicate (m + n) a = replicate m a + replicate n a := congr_arg _ <| List.replicate_add .. -/-- `Multiset.replicate` as an `AddMonoidHom`. -/ -@[simps] -def replicateAddMonoidHom (a : α) : ℕ →+ Multiset α where - toFun := fun n => replicate n a - map_zero' := replicate_zero a - map_add' := fun _ _ => replicate_add _ _ a - theorem replicate_one (a : α) : replicate 1 a = {a} := rfl @[simp] theorem card_replicate (n) (a : α) : card (replicate n a) = n := @@ -881,12 +850,6 @@ theorem replicate_subset_singleton (n : ℕ) (a : α) : replicate n a ⊆ {a} := theorem replicate_le_coe {a : α} {n} {l : List α} : replicate n a ≤ l ↔ List.replicate n a <+ l := ⟨fun ⟨_l', p, s⟩ => perm_replicate.1 p ▸ s, Sublist.subperm⟩ -theorem nsmul_replicate {a : α} (n m : ℕ) : n • replicate m a = replicate (n * m) a := - ((replicateAddMonoidHom a).map_nsmul _ _).symm - -theorem nsmul_singleton (a : α) (n) : n • ({a} : Multiset α) = replicate n a := by - rw [← replicate_one, nsmul_replicate, mul_one] - theorem replicate_le_replicate (a : α) {k n : ℕ} : replicate k a ≤ replicate n a ↔ k ≤ n := _root_.trans (by rw [← replicate_le_coe, coe_replicate]) (List.replicate_sublist_replicate a) @@ -961,7 +924,8 @@ theorem le_cons_erase (s : Multiset α) (a : α) : s ≤ a ::ₘ s.erase a := else by rw [erase_of_not_mem h]; apply le_cons_self theorem add_singleton_eq_iff {s t : Multiset α} {a : α} : s + {a} = t ↔ a ∈ t ∧ s = t.erase a := by - rw [add_comm, singleton_add]; constructor + rw [Multiset.add_comm, singleton_add] + constructor · rintro rfl exact ⟨s.mem_cons_self a, (s.erase_cons_head a).symm⟩ · rintro ⟨h, rfl⟩ @@ -970,15 +934,15 @@ theorem add_singleton_eq_iff {s t : Multiset α} {a : α} : s + {a} = t ↔ a theorem erase_add_left_pos {a : α} {s : Multiset α} (t) : a ∈ s → (s + t).erase a = s.erase a + t := Quotient.inductionOn₂ s t fun _l₁ l₂ h => congr_arg _ <| erase_append_left l₂ h -theorem erase_add_right_pos {a : α} (s) {t : Multiset α} (h : a ∈ t) : - (s + t).erase a = s + t.erase a := by rw [add_comm, erase_add_left_pos s h, add_comm] +theorem erase_add_right_pos {a : α} (s) (h : a ∈ t) : (s + t).erase a = s + t.erase a := by + rw [Multiset.add_comm, erase_add_left_pos s h, Multiset.add_comm] theorem erase_add_right_neg {a : α} {s : Multiset α} (t) : a ∉ s → (s + t).erase a = s + t.erase a := Quotient.inductionOn₂ s t fun _l₁ l₂ h => congr_arg _ <| erase_append_right l₂ h -theorem erase_add_left_neg {a : α} (s) {t : Multiset α} (h : a ∉ t) : - (s + t).erase a = s.erase a + t := by rw [add_comm, erase_add_right_neg s h, add_comm] +theorem erase_add_left_neg {a : α} (s) (h : a ∉ t) : (s + t).erase a = s.erase a + t := by + rw [Multiset.add_comm, erase_add_right_neg s h, Multiset.add_comm] theorem erase_le (a : α) (s : Multiset α) : s.erase a ≤ s := Quot.inductionOn s fun l => (erase_sublist a l).subperm @@ -1097,20 +1061,6 @@ instance canLift (c) (p) [CanLift α β c p] : lift l to List β using hl exact ⟨l, map_coe _ _⟩ -/-- `Multiset.map` as an `AddMonoidHom`. -/ -def mapAddMonoidHom (f : α → β) : Multiset α →+ Multiset β where - toFun := map f - map_zero' := map_zero _ - map_add' := map_add _ - -@[simp] -theorem coe_mapAddMonoidHom (f : α → β) : - (mapAddMonoidHom f : Multiset α → Multiset β) = map f := - rfl - -theorem map_nsmul (f : α → β) (n : ℕ) (s) : map f (n • s) = n • map f s := - (mapAddMonoidHom f).map_nsmul _ _ - @[simp] theorem mem_map {f : α → β} {b : β} {s : Multiset α} : b ∈ map f s ↔ ∃ a, a ∈ s ∧ f a = b := Quot.inductionOn s fun _l => List.mem_map @@ -1463,260 +1413,6 @@ instance decidableDexistsMultiset {p : ∀ a ∈ m, Prop} [_hp : ∀ (a) (h : a end DecidablePiExists -/-! ### Subtraction -/ - - -section - -variable [DecidableEq α] {s t u : Multiset α} {a : α} - -/-- `s - t` is the multiset such that `count a (s - t) = count a s - count a t` for all `a` - (note that it is truncated subtraction, so it is `0` if `count a t ≥ count a s`). -/ -protected def sub (s t : Multiset α) : Multiset α := - (Quotient.liftOn₂ s t fun l₁ l₂ => (l₁.diff l₂ : Multiset α)) fun _v₁ _v₂ _w₁ _w₂ p₁ p₂ => - Quot.sound <| p₁.diff p₂ - -instance : Sub (Multiset α) := - ⟨Multiset.sub⟩ - -@[simp] -theorem coe_sub (s t : List α) : (s - t : Multiset α) = (s.diff t : List α) := - rfl - -/-- This is a special case of `tsub_zero`, which should be used instead of this. - This is needed to prove `OrderedSub (Multiset α)`. -/ -protected theorem sub_zero (s : Multiset α) : s - 0 = s := - Quot.inductionOn s fun _l => rfl - -@[simp] -theorem sub_cons (a : α) (s t : Multiset α) : s - a ::ₘ t = s.erase a - t := - Quotient.inductionOn₂ s t fun _l₁ _l₂ => congr_arg _ <| diff_cons _ _ _ - -protected theorem zero_sub (t : Multiset α) : 0 - t = 0 := - Multiset.induction_on t rfl fun a s ih => by simp [ih] - -/-- This is a special case of `tsub_le_iff_right`, which should be used instead of this. - This is needed to prove `OrderedSub (Multiset α)`. -/ -protected theorem sub_le_iff_le_add : s - t ≤ u ↔ s ≤ u + t := by - revert s - exact @(Multiset.induction_on t (by simp [Multiset.sub_zero]) fun a t IH s => by - simp [IH, erase_le_iff_le_cons]) - -protected theorem sub_le_self (s t : Multiset α) : s - t ≤ s := by - rw [Multiset.sub_le_iff_le_add] - exact le_add_right _ _ - -instance : OrderedSub (Multiset α) := - ⟨fun _n _m _k => Multiset.sub_le_iff_le_add⟩ - -instance : ExistsAddOfLE (Multiset α) where - exists_add_of_le h := leInductionOn h fun s => - let ⟨l, p⟩ := s.exists_perm_append - ⟨l, Quot.sound p⟩ - -theorem cons_sub_of_le (a : α) {s t : Multiset α} (h : t ≤ s) : a ::ₘ s - t = a ::ₘ (s - t) := by - rw [← singleton_add, ← singleton_add, add_tsub_assoc_of_le h] - -theorem sub_eq_fold_erase (s t : Multiset α) : s - t = foldl erase s t := - Quotient.inductionOn₂ s t fun l₁ l₂ => by - show ofList (l₁.diff l₂) = foldl erase l₁ l₂ - rw [diff_eq_foldl l₁ l₂] - symm - exact foldl_hom _ _ _ _ _ fun x y => rfl - -@[simp] -theorem card_sub {s t : Multiset α} (h : t ≤ s) : card (s - t) = card s - card t := - Nat.eq_sub_of_add_eq <| by rw [← card_add, tsub_add_cancel_of_le h] - -/-! ### Union -/ - - -/-- `s ∪ t` is the lattice join operation with respect to the - multiset `≤`. The multiplicity of `a` in `s ∪ t` is the maximum - of the multiplicities in `s` and `t`. -/ -def union (s t : Multiset α) : Multiset α := - s - t + t - -instance : Union (Multiset α) := - ⟨union⟩ - -theorem union_def (s t : Multiset α) : s ∪ t = s - t + t := - rfl - -theorem le_union_left {s t : Multiset α} : s ≤ s ∪ t := - le_tsub_add - -theorem le_union_right {s t : Multiset α} : t ≤ s ∪ t := - le_add_left _ _ - -theorem eq_union_left : t ≤ s → s ∪ t = s := - tsub_add_cancel_of_le - -@[gcongr] -theorem union_le_union_right (h : s ≤ t) (u) : s ∪ u ≤ t ∪ u := - add_le_add_right (tsub_le_tsub_right h _) u - -theorem union_le (h₁ : s ≤ u) (h₂ : t ≤ u) : s ∪ t ≤ u := by - rw [← eq_union_left h₂]; exact union_le_union_right h₁ t - - -@[simp] -theorem mem_union : a ∈ s ∪ t ↔ a ∈ s ∨ a ∈ t := - ⟨fun h => (mem_add.1 h).imp_left (mem_of_le <| Multiset.sub_le_self _ _), - (Or.elim · (mem_of_le le_union_left) (mem_of_le le_union_right))⟩ - -@[simp] -theorem map_union [DecidableEq β] {f : α → β} (finj : Function.Injective f) {s t : Multiset α} : - map f (s ∪ t) = map f s ∪ map f t := - Quotient.inductionOn₂ s t fun l₁ l₂ => - congr_arg ofList (by rw [List.map_append f, List.map_diff finj]) - -@[simp] theorem zero_union : 0 ∪ s = s := by - simp [union_def, Multiset.zero_sub] - -@[simp] theorem union_zero : s ∪ 0 = s := by - simp [union_def] - -/-! ### Intersection -/ - -/-- `s ∩ t` is the lattice meet operation with respect to the - multiset `≤`. The multiplicity of `a` in `s ∩ t` is the minimum - of the multiplicities in `s` and `t`. -/ -def inter (s t : Multiset α) : Multiset α := - Quotient.liftOn₂ s t (fun l₁ l₂ => (l₁.bagInter l₂ : Multiset α)) fun _v₁ _v₂ _w₁ _w₂ p₁ p₂ => - Quot.sound <| p₁.bagInter p₂ - -instance : Inter (Multiset α) := - ⟨inter⟩ - -@[simp] -theorem inter_zero (s : Multiset α) : s ∩ 0 = 0 := - Quot.inductionOn s fun l => congr_arg ofList l.bagInter_nil - -@[simp] -theorem zero_inter (s : Multiset α) : 0 ∩ s = 0 := - Quot.inductionOn s fun l => congr_arg ofList l.nil_bagInter - -@[simp] -theorem cons_inter_of_pos {a} (s : Multiset α) {t} : a ∈ t → (a ::ₘ s) ∩ t = a ::ₘ s ∩ t.erase a := - Quotient.inductionOn₂ s t fun _l₁ _l₂ h => congr_arg ofList <| cons_bagInter_of_pos _ h - -@[simp] -theorem cons_inter_of_neg {a} (s : Multiset α) {t} : a ∉ t → (a ::ₘ s) ∩ t = s ∩ t := - Quotient.inductionOn₂ s t fun _l₁ _l₂ h => congr_arg ofList <| cons_bagInter_of_neg _ h - -theorem inter_le_left {s t : Multiset α} : s ∩ t ≤ s := - Quotient.inductionOn₂ s t fun _l₁ _l₂ => (bagInter_sublist_left _ _).subperm - -theorem inter_le_right {s t : Multiset α} : s ∩ t ≤ t := by - induction' s using Multiset.induction_on with a s IH generalizing t - · exact (zero_inter t).symm ▸ zero_le _ - by_cases h : a ∈ t - · simpa [h] using cons_le_cons a (IH (t := t.erase a)) - · simp [h, IH] - -theorem le_inter (h₁ : s ≤ t) (h₂ : s ≤ u) : s ≤ t ∩ u := by - revert s u; refine @(Multiset.induction_on t ?_ fun a t IH => ?_) <;> intros s u h₁ h₂ - · simpa only [zero_inter] using h₁ - by_cases h : a ∈ u - · rw [cons_inter_of_pos _ h, ← erase_le_iff_le_cons] - exact IH (erase_le_iff_le_cons.2 h₁) (erase_le_erase _ h₂) - · rw [cons_inter_of_neg _ h] - exact IH ((le_cons_of_not_mem <| mt (mem_of_le h₂) h).1 h₁) h₂ - -@[simp] -theorem mem_inter : a ∈ s ∩ t ↔ a ∈ s ∧ a ∈ t := - ⟨fun h => ⟨mem_of_le inter_le_left h, mem_of_le inter_le_right h⟩, fun ⟨h₁, h₂⟩ => by - rw [← cons_erase h₁, cons_inter_of_pos _ h₂]; apply mem_cons_self⟩ - -instance : Lattice (Multiset α) where - sup := (· ∪ ·) - sup_le _ _ _ := union_le - le_sup_left _ _ := le_union_left - le_sup_right _ _ := le_union_right - inf := (· ∩ ·) - le_inf _ _ _ := le_inter - inf_le_left _ _ := inter_le_left - inf_le_right _ _ := inter_le_right - -@[simp] -theorem sup_eq_union (s t : Multiset α) : s ⊔ t = s ∪ t := - rfl - -@[simp] -theorem inf_eq_inter (s t : Multiset α) : s ⊓ t = s ∩ t := - rfl - -@[simp] -theorem le_inter_iff : s ≤ t ∩ u ↔ s ≤ t ∧ s ≤ u := - le_inf_iff - -@[simp] -theorem union_le_iff : s ∪ t ≤ u ↔ s ≤ u ∧ t ≤ u := - sup_le_iff - -theorem union_comm (s t : Multiset α) : s ∪ t = t ∪ s := sup_comm _ _ - -theorem inter_comm (s t : Multiset α) : s ∩ t = t ∩ s := inf_comm _ _ - -theorem eq_union_right (h : s ≤ t) : s ∪ t = t := by rw [union_comm, eq_union_left h] - -@[gcongr] -theorem union_le_union_left (h : s ≤ t) (u) : u ∪ s ≤ u ∪ t := - sup_le_sup_left h _ - -theorem union_le_add (s t : Multiset α) : s ∪ t ≤ s + t := - union_le (le_add_right _ _) (le_add_left _ _) - -theorem union_add_distrib (s t u : Multiset α) : s ∪ t + u = s + u ∪ (t + u) := by - simpa [(· ∪ ·), union, eq_comm, add_assoc] using - show s + u - (t + u) = s - t by rw [add_comm t, tsub_add_eq_tsub_tsub, add_tsub_cancel_right] - -theorem add_union_distrib (s t u : Multiset α) : s + (t ∪ u) = s + t ∪ (s + u) := by - rw [add_comm, union_add_distrib, add_comm s, add_comm s] - -theorem cons_union_distrib (a : α) (s t : Multiset α) : a ::ₘ (s ∪ t) = a ::ₘ s ∪ a ::ₘ t := by - simpa using add_union_distrib (a ::ₘ 0) s t - -theorem inter_add_distrib (s t u : Multiset α) : s ∩ t + u = (s + u) ∩ (t + u) := by - by_contra! h - obtain ⟨a, ha⟩ := lt_iff_cons_le.1 <| h.lt_of_le <| le_inter - (add_le_add_right inter_le_left _) (add_le_add_right inter_le_right _) - rw [← cons_add] at ha - exact (lt_cons_self (s ∩ t) a).not_le <| le_inter - (le_of_add_le_add_right (ha.trans inter_le_left)) - (le_of_add_le_add_right (ha.trans inter_le_right)) - -theorem add_inter_distrib (s t u : Multiset α) : s + t ∩ u = (s + t) ∩ (s + u) := by - rw [add_comm, inter_add_distrib, add_comm s, add_comm s] - -theorem cons_inter_distrib (a : α) (s t : Multiset α) : a ::ₘ s ∩ t = (a ::ₘ s) ∩ (a ::ₘ t) := by - simp - -lemma union_add_inter (s t : Multiset α) : s ∪ t + s ∩ t = s + t := by - apply _root_.le_antisymm - · rw [union_add_distrib] - refine union_le (add_le_add_left inter_le_right _) ?_ - rw [add_comm] - exact add_le_add_right inter_le_left _ - · rw [add_comm, add_inter_distrib] - refine le_inter (add_le_add_right le_union_right _) ?_ - rw [add_comm] - exact add_le_add_right le_union_left _ - -theorem sub_add_inter (s t : Multiset α) : s - t + s ∩ t = s := by - rw [inter_comm] - revert s; refine Multiset.induction_on t (by simp) fun a t IH s => ?_ - by_cases h : a ∈ s - · rw [cons_inter_of_pos _ h, sub_cons, add_cons, IH, cons_erase h] - · rw [cons_inter_of_neg _ h, sub_cons, erase_of_not_mem h, IH] - -theorem sub_inter (s t : Multiset α) : s - s ∩ t = s - t := - add_right_cancel (b := s ∩ t) <| by - rw [sub_add_inter s t, tsub_add_cancel_of_le inter_le_left] - -end - /-! ### `Multiset.filter` -/ @@ -1803,54 +1499,14 @@ theorem filter_cons {a : α} (s : Multiset α) : filter p (a ::ₘ s) = (if p a then {a} else 0) + filter p s := by split_ifs with h · rw [filter_cons_of_pos _ h, singleton_add] - · rw [filter_cons_of_neg _ h, zero_add] + · rw [filter_cons_of_neg _ h, Multiset.zero_add] theorem filter_singleton {a : α} (p : α → Prop) [DecidablePred p] : filter p {a} = if p a then {a} else ∅ := by - simp only [singleton, filter_cons, filter_zero, add_zero, empty_eq_zero] - -theorem filter_nsmul (s : Multiset α) (n : ℕ) : filter p (n • s) = n • filter p s := by - refine s.induction_on ?_ ?_ - · simp only [filter_zero, nsmul_zero] - · intro a ha ih - rw [nsmul_cons, filter_add, ih, filter_cons, nsmul_add] - congr - split_ifs with hp <;> - · simp only [filter_eq_self, nsmul_zero, filter_eq_nil] - intro b hb - rwa [mem_singleton.mp (mem_of_mem_nsmul hb)] + simp only [singleton, filter_cons, filter_zero, Multiset.add_zero, empty_eq_zero] variable (p) -@[simp] -theorem filter_sub [DecidableEq α] (s t : Multiset α) : - filter p (s - t) = filter p s - filter p t := by - revert s; refine Multiset.induction_on t (by simp) fun a t IH s => ?_ - rw [sub_cons, IH] - by_cases h : p a - · rw [filter_cons_of_pos _ h, sub_cons] - congr - by_cases m : a ∈ s - · rw [← cons_inj_right a, ← filter_cons_of_pos _ h, cons_erase (mem_filter_of_mem m h), - cons_erase m] - · rw [erase_of_not_mem m, erase_of_not_mem (mt mem_of_mem_filter m)] - · rw [filter_cons_of_neg _ h] - by_cases m : a ∈ s - · rw [(by rw [filter_cons_of_neg _ h] : filter p (erase s a) = filter p (a ::ₘ erase s a)), - cons_erase m] - · rw [erase_of_not_mem m] - -@[simp] -theorem filter_union [DecidableEq α] (s t : Multiset α) : - filter p (s ∪ t) = filter p s ∪ filter p t := by simp [(· ∪ ·), union] - -@[simp] -theorem filter_inter [DecidableEq α] (s t : Multiset α) : - filter p (s ∩ t) = filter p s ∩ filter p t := - le_antisymm (le_inter (filter_le_filter _ inter_le_left) (filter_le_filter _ inter_le_right)) <| - le_filter.2 ⟨inf_le_inf (filter_le _ _) (filter_le _ _), fun _a h => - of_mem_filter (mem_of_le inter_le_left h)⟩ - @[simp] theorem filter_filter (q) [DecidablePred q] (s : Multiset α) : filter p (filter q s) = filter (fun a => p a ∧ q a) s := @@ -1865,7 +1521,7 @@ theorem filter_add_filter (q) [DecidablePred q] (s : Multiset α) : theorem filter_add_not (s : Multiset α) : filter p s + filter (fun a => ¬p a) s = s := by rw [filter_add_filter, filter_eq_self.2, filter_eq_nil.2] - · simp only [add_zero] + · simp only [Multiset.add_zero] · simp [Decidable.em, -Bool.not_eq_true, -not_and, not_and_or, or_comm] · simp only [Bool.not_eq_true, decide_eq_true_eq, Bool.eq_false_or_eq_true, decide_true, implies_true, Decidable.em] @@ -2008,29 +1664,9 @@ theorem countP_le_card (s) : countP p s ≤ card s := theorem countP_add (s t) : countP p (s + t) = countP p s + countP p t := by simp [countP_eq_card_filter] -@[simp] -theorem countP_nsmul (s) (n : ℕ) : countP p (n • s) = n * countP p s := by - induction n <;> simp [*, succ_nsmul, succ_mul, zero_nsmul] - theorem card_eq_countP_add_countP (s) : card s = countP p s + countP (fun x => ¬p x) s := Quot.inductionOn s fun l => by simp [l.length_eq_countP_add_countP p] -/-- `countP p`, the number of elements of a multiset satisfying `p`, promoted to an -`AddMonoidHom`. -/ -def countPAddMonoidHom : Multiset α →+ ℕ where - toFun := countP p - map_zero' := countP_zero _ - map_add' := countP_add _ - -@[simp] -theorem coe_countPAddMonoidHom : (countPAddMonoidHom p : Multiset α → ℕ) = countP p := - rfl - -@[simp] -theorem countP_sub [DecidableEq α] {s t : Multiset α} (h : t ≤ s) : - countP p (s - t) = countP p s - countP p t := by - simp [countP_eq_card_filter, h, filter_le_filter] - @[gcongr] theorem countP_le_of_le {s t} (h : s ≤ t) : countP p s ≤ countP p t := by simpa [countP_eq_card_filter] using card_le_card (filter_le_filter p h) @@ -2058,7 +1694,7 @@ theorem countP_map (f : α → β) (s : Multiset α) (p : β → Prop) [Decidabl refine Multiset.induction_on s ?_ fun a t IH => ?_ · rw [map_zero, countP_zero, filter_zero, card_zero] · rw [map_cons, countP_cons, IH, filter_cons, card_add, apply_ite card, card_zero, card_singleton, - add_comm] + Nat.add_comm] -- Porting note: `Lean.Internal.coeM` forces us to type-ascript `{a // a ∈ s}` lemma countP_attach (s : Multiset α) : s.attach.countP (fun a : {a // a ∈ s} ↦ p a) = s.countP p := @@ -2109,7 +1745,7 @@ end section -variable [DecidableEq α] {s : Multiset α} +variable [DecidableEq α] {s t u : Multiset α} /-- `count a s` is the multiplicity of `a` in `s`. -/ def count (a : α) : Multiset α → ℕ := @@ -2150,24 +1786,12 @@ theorem count_singleton_self (a : α) : count a ({a} : Multiset α) = 1 := count_eq_one_of_mem (nodup_singleton a) <| mem_singleton_self a theorem count_singleton (a b : α) : count a ({b} : Multiset α) = if a = b then 1 else 0 := by - simp only [count_cons, ← cons_zero, count_zero, zero_add] + simp only [count_cons, ← cons_zero, count_zero, Nat.zero_add] @[simp] theorem count_add (a : α) : ∀ s t, count a (s + t) = count a s + count a t := countP_add _ -/-- `count a`, the multiplicity of `a` in a multiset, promoted to an `AddMonoidHom`. -/ -def countAddMonoidHom (a : α) : Multiset α →+ ℕ := - countPAddMonoidHom (a = ·) - -@[simp] -theorem coe_countAddMonoidHom {a : α} : (countAddMonoidHom a : Multiset α → ℕ) = count a := - rfl - -@[simp] -theorem count_nsmul (a : α) (n s) : count a (n • s) = n * count a s := by - induction n <;> simp [*, succ_nsmul, succ_mul, zero_nsmul] - @[simp] lemma count_attach (a : {x // x ∈ s}) : s.attach.count a = s.count ↑a := Eq.trans (countP_congr rfl fun _ _ => by simp [Subtype.ext_iff]) <| countP_attach _ _ @@ -2209,23 +1833,6 @@ theorem count_erase_of_ne {a b : α} (ab : a ≠ b) (s : Multiset α) : Quotient.inductionOn s fun l => by convert List.count_erase_of_ne ab l <;> rw [← coe_count] <;> simp -@[simp] -theorem count_sub (a : α) (s t : Multiset α) : count a (s - t) = count a s - count a t := by - revert s; refine Multiset.induction_on t (by simp) fun b t IH s => ?_ - rw [sub_cons, IH] - rcases Decidable.eq_or_ne a b with rfl | ab - · rw [count_erase_self, count_cons_self, Nat.sub_sub, add_comm] - · rw [count_erase_of_ne ab, count_cons_of_ne ab] - -@[simp] -theorem count_union (a : α) (s t : Multiset α) : count a (s ∪ t) = max (count a s) (count a t) := by - simp [(· ∪ ·), union, Nat.sub_add_eq_max] - -@[simp] -theorem count_inter (a : α) (s t : Multiset α) : count a (s ∩ t) = min (count a s) (count a t) := by - apply @Nat.add_left_cancel (count a (s - t)) - rw [← count_add, sub_add_inter, count_sub, Nat.sub_add_min_cancel] - theorem le_count_iff_replicate_le {a : α} {s : Multiset α} {n : ℕ} : n ≤ count a s ↔ replicate n a ≤ s := Quot.inductionOn s fun _l => by @@ -2263,20 +1870,8 @@ theorem ext' {s t : Multiset α} : (∀ a, count a s = count a t) → s = t := lemma count_injective : Injective fun (s : Multiset α) a ↦ s.count a := fun _s _t hst ↦ ext' <| congr_fun hst -@[simp] -theorem coe_inter (s t : List α) : (s ∩ t : Multiset α) = (s.bagInter t : List α) := by ext; simp - theorem le_iff_count {s t : Multiset α} : s ≤ t ↔ ∀ a, count a s ≤ count a t := - ⟨fun h a => count_le_of_le a h, fun al => by - rw [← (ext.2 fun a => by simp [max_eq_right (al a)] : s ∪ t = t)]; apply le_union_left⟩ - -instance : DistribLattice (Multiset α) := - { le_sup_inf := fun s t u => - le_of_eq <| - Eq.symm <| - ext.2 fun a => by - simp only [max_min_distrib_left, Multiset.count_inter, Multiset.sup_eq_union, - Multiset.count_union, Multiset.inf_eq_inter] } + Quotient.inductionOn₂ s t fun _ _ ↦ by simp [subperm_iff_count] theorem count_map {α β : Type*} (f : α → β) (s : Multiset α) [DecidableEq β] (b : β) : count b (map f s) = card (s.filter fun a => b = f a) := by @@ -2303,11 +1898,6 @@ theorem count_map_eq_count' [DecidableEq β] (f : α → β) (s : Multiset α) ( rw [hf hkx] at hks contradiction -@[simp] -theorem sub_filter_eq_filter_not (p) [DecidablePred p] (s : Multiset α) : - s - s.filter p = s.filter (fun a ↦ ¬ p a) := by - ext a; by_cases h : p a <;> simp [h] - theorem filter_eq' (s : Multiset α) (b : α) : s.filter (· = b) = replicate (count b s) b := Quotient.inductionOn s fun l => by simp only [quot_mk_to_coe, filter_coe, mem_coe, coe_count] @@ -2316,6 +1906,317 @@ theorem filter_eq' (s : Multiset α) (b : α) : s.filter (· = b) = replicate (c theorem filter_eq (s : Multiset α) (b : α) : s.filter (Eq b) = replicate (count b s) b := by simp_rw [← filter_eq', eq_comm] +lemma erase_attach_map_val (s : Multiset α) (x : {x // x ∈ s}) : + (s.attach.erase x).map (↑) = s.erase x := by + rw [Multiset.map_erase _ val_injective, attach_map_val] + +lemma erase_attach_map (s : Multiset α) (f : α → β) (x : {x // x ∈ s}) : + (s.attach.erase x).map (fun j : {x // x ∈ s} ↦ f j) = (s.erase x).map f := by + simp only [← Function.comp_apply (f := f)] + rw [← map_map, erase_attach_map_val] + +omit [DecidableEq α] in +protected lemma add_left_inj : s + u = t + u ↔ s = t := by classical simp [Multiset.ext] + +omit [DecidableEq α] in +protected lemma add_right_inj : s + t = s + u ↔ t = u := by classical simp [Multiset.ext] + +end + +/-! ### Subtraction -/ + +section sub +variable [DecidableEq α] {s t u : Multiset α} {a : α} + +/-- `s - t` is the multiset such that `count a (s - t) = count a s - count a t` for all `a`. +(note that it is truncated subtraction, so `count a (s - t) = 0` if `count a s ≤ count a t`). -/ +protected def sub (s t : Multiset α) : Multiset α := + (Quotient.liftOn₂ s t fun l₁ l₂ => (l₁.diff l₂ : Multiset α)) fun _v₁ _v₂ _w₁ _w₂ p₁ p₂ => + Quot.sound <| p₁.diff p₂ + +instance : Sub (Multiset α) := ⟨.sub⟩ + +@[simp] +lemma coe_sub (s t : List α) : (s - t : Multiset α) = s.diff t := + rfl + +/-- This is a special case of `tsub_zero`, which should be used instead of this. +This is needed to prove `OrderedSub (Multiset α)`. -/ +@[simp, nolint simpNF] -- We want to use this lemma earlier than the lemma simp can prove it with +protected lemma sub_zero (s : Multiset α) : s - 0 = s := + Quot.inductionOn s fun _l => rfl + +@[simp] +lemma sub_cons (a : α) (s t : Multiset α) : s - a ::ₘ t = s.erase a - t := + Quotient.inductionOn₂ s t fun _l₁ _l₂ => congr_arg _ <| diff_cons _ _ _ + +protected lemma zero_sub (t : Multiset α) : 0 - t = 0 := + Multiset.induction_on t rfl fun a s ih => by simp [ih] + +@[simp] +lemma countP_sub {s t : Multiset α} : + t ≤ s → ∀ (p : α → Prop) [DecidablePred p], countP p (s - t) = countP p s - countP p t := + Quotient.inductionOn₂ s t fun _l₁ _l₂ hl _ _ ↦ List.countP_diff hl _ + +@[simp] +lemma count_sub (a : α) (s t : Multiset α) : count a (s - t) = count a s - count a t := + Quotient.inductionOn₂ s t <| by simp [List.count_diff] + +/-- This is a special case of `tsub_le_iff_right`, which should be used instead of this. +This is needed to prove `OrderedSub (Multiset α)`. -/ +protected lemma sub_le_iff_le_add : s - t ≤ u ↔ s ≤ u + t := by + revert s + exact @(Multiset.induction_on t (by simp [Multiset.sub_zero]) fun a t IH s => by + simp [IH, erase_le_iff_le_cons]) + +/-- This is a special case of `tsub_le_iff_left`, which should be used instead of this. -/ +protected lemma sub_le_iff_le_add' : s - t ≤ u ↔ s ≤ t + u := by + rw [Multiset.sub_le_iff_le_add, Multiset.add_comm] + +protected theorem sub_le_self (s t : Multiset α) : s - t ≤ s := by + rw [Multiset.sub_le_iff_le_add] + exact le_add_right _ _ + +protected lemma add_sub_assoc (hut : u ≤ t) : s + t - u = s + (t - u) := by + ext a; simp [Nat.add_sub_assoc <| count_le_of_le _ hut] + +protected lemma add_sub_cancel (hts : t ≤ s) : s - t + t = s := by + ext a; simp [Nat.sub_add_cancel <| count_le_of_le _ hts] + +protected lemma sub_add_cancel (hts : t ≤ s) : s - t + t = s := by + ext a; simp [Nat.sub_add_cancel <| count_le_of_le _ hts] + +protected lemma sub_add_eq_sub_sub : s - (t + u) = s - t - u := by ext; simp [Nat.sub_add_eq] + +protected lemma le_sub_add : s ≤ s - t + t := Multiset.sub_le_iff_le_add.1 le_rfl +protected lemma le_add_sub : s ≤ t + (s - t) := Multiset.sub_le_iff_le_add'.1 le_rfl + +protected lemma sub_le_sub_right (hst : s ≤ t) : s - u ≤ t - u := + Multiset.sub_le_iff_le_add'.mpr <| hst.trans Multiset.le_add_sub + +protected lemma add_sub_cancel_right : s + t - t = s := by ext a; simp + +protected lemma eq_sub_of_add_eq (hstu : s + t = u) : s = u - t := by + rw [← hstu, Multiset.add_sub_cancel_right] + +@[simp] +lemma filter_sub (p : α → Prop) [DecidablePred p] (s t : Multiset α) : + filter p (s - t) = filter p s - filter p t := by + revert s; refine Multiset.induction_on t (by simp) fun a t IH s => ?_ + rw [sub_cons, IH] + by_cases h : p a + · rw [filter_cons_of_pos _ h, sub_cons] + congr + by_cases m : a ∈ s + · rw [← cons_inj_right a, ← filter_cons_of_pos _ h, cons_erase (mem_filter_of_mem m h), + cons_erase m] + · rw [erase_of_not_mem m, erase_of_not_mem (mt mem_of_mem_filter m)] + · rw [filter_cons_of_neg _ h] + by_cases m : a ∈ s + · rw [(by rw [filter_cons_of_neg _ h] : filter p (erase s a) = filter p (a ::ₘ erase s a)), + cons_erase m] + · rw [erase_of_not_mem m] + +@[simp] +lemma sub_filter_eq_filter_not (p : α → Prop) [DecidablePred p] (s : Multiset α) : + s - s.filter p = s.filter fun a ↦ ¬ p a := by ext a; by_cases h : p a <;> simp [h] + +lemma cons_sub_of_le (a : α) {s t : Multiset α} (h : t ≤ s) : a ::ₘ s - t = a ::ₘ (s - t) := by + rw [← singleton_add, ← singleton_add, Multiset.add_sub_assoc h] + +lemma sub_eq_fold_erase (s t : Multiset α) : s - t = foldl erase s t := + Quotient.inductionOn₂ s t fun l₁ l₂ => by + show ofList (l₁.diff l₂) = foldl erase l₁ l₂ + rw [diff_eq_foldl l₁ l₂] + symm + exact foldl_hom _ _ _ _ _ fun x y => rfl + +@[simp] +lemma card_sub {s t : Multiset α} (h : t ≤ s) : card (s - t) = card s - card t := + Nat.eq_sub_of_add_eq <| by rw [← card_add, Multiset.sub_add_cancel h] + +/-! ### Union -/ + +/-- `s ∪ t` is the multiset such that the multiplicity of each `a` in it is the maximum of the +multiplicity of `a` in `s` and `t`. This is the supremum of multisets. -/ +def union (s t : Multiset α) : Multiset α := s - t + t + +instance : Union (Multiset α) := + ⟨union⟩ + +lemma union_def (s t : Multiset α) : s ∪ t = s - t + t := rfl + +lemma le_union_left : s ≤ s ∪ t := Multiset.le_sub_add +lemma le_union_right : t ≤ s ∪ t := le_add_left _ _ +lemma eq_union_left : t ≤ s → s ∪ t = s := Multiset.sub_add_cancel + +@[gcongr] +lemma union_le_union_right (h : s ≤ t) (u) : s ∪ u ≤ t ∪ u := + Multiset.add_le_add_right <| Multiset.sub_le_sub_right h + +lemma union_le (h₁ : s ≤ u) (h₂ : t ≤ u) : s ∪ t ≤ u := by + rw [← eq_union_left h₂]; exact union_le_union_right h₁ t + +@[simp] +lemma mem_union : a ∈ s ∪ t ↔ a ∈ s ∨ a ∈ t := + ⟨fun h => (mem_add.1 h).imp_left (mem_of_le <| Multiset.sub_le_self _ _), + (Or.elim · (mem_of_le le_union_left) (mem_of_le le_union_right))⟩ + +@[simp] +lemma map_union [DecidableEq β] {f : α → β} (finj : Function.Injective f) {s t : Multiset α} : + map f (s ∪ t) = map f s ∪ map f t := + Quotient.inductionOn₂ s t fun l₁ l₂ => + congr_arg ofList (by rw [List.map_append f, List.map_diff finj]) + +@[simp] lemma zero_union : 0 ∪ s = s := by simp [union_def, Multiset.zero_sub] +@[simp] lemma union_zero : s ∪ 0 = s := by simp [union_def] + +@[simp] +lemma count_union (a : α) (s t : Multiset α) : count a (s ∪ t) = max (count a s) (count a t) := by + simp [(· ∪ ·), union, Nat.sub_add_eq_max] + +@[simp] lemma filter_union (p : α → Prop) [DecidablePred p] (s t : Multiset α) : + filter p (s ∪ t) = filter p s ∪ filter p t := by simp [(· ∪ ·), union] + +/-! ### Intersection -/ + +/-- `s ∩ t` is the multiset such that the multiplicity of each `a` in it is the minimum of the +multiplicity of `a` in `s` and `t`. This is the infimum of multisets. -/ +def inter (s t : Multiset α) : Multiset α := + Quotient.liftOn₂ s t (fun l₁ l₂ => (l₁.bagInter l₂ : Multiset α)) fun _v₁ _v₂ _w₁ _w₂ p₁ p₂ => + Quot.sound <| p₁.bagInter p₂ + +instance : Inter (Multiset α) := ⟨inter⟩ + +@[simp] lemma inter_zero (s : Multiset α) : s ∩ 0 = 0 := + Quot.inductionOn s fun l => congr_arg ofList l.bagInter_nil + +@[simp] lemma zero_inter (s : Multiset α) : 0 ∩ s = 0 := + Quot.inductionOn s fun l => congr_arg ofList l.nil_bagInter + +@[simp] +lemma cons_inter_of_pos (s : Multiset α) : a ∈ t → (a ::ₘ s) ∩ t = a ::ₘ s ∩ t.erase a := + Quotient.inductionOn₂ s t fun _l₁ _l₂ h => congr_arg ofList <| cons_bagInter_of_pos _ h + +@[simp] +lemma cons_inter_of_neg (s : Multiset α) : a ∉ t → (a ::ₘ s) ∩ t = s ∩ t := + Quotient.inductionOn₂ s t fun _l₁ _l₂ h => congr_arg ofList <| cons_bagInter_of_neg _ h + +lemma inter_le_left : s ∩ t ≤ s := + Quotient.inductionOn₂ s t fun _l₁ _l₂ => (bagInter_sublist_left _ _).subperm + +lemma inter_le_right : s ∩ t ≤ t := by + induction' s using Multiset.induction_on with a s IH generalizing t + · exact (zero_inter t).symm ▸ zero_le _ + by_cases h : a ∈ t + · simpa [h] using cons_le_cons a (IH (t := t.erase a)) + · simp [h, IH] + +lemma le_inter (h₁ : s ≤ t) (h₂ : s ≤ u) : s ≤ t ∩ u := by + revert s u; refine @(Multiset.induction_on t ?_ fun a t IH => ?_) <;> intros s u h₁ h₂ + · simpa only [zero_inter] using h₁ + by_cases h : a ∈ u + · rw [cons_inter_of_pos _ h, ← erase_le_iff_le_cons] + exact IH (erase_le_iff_le_cons.2 h₁) (erase_le_erase _ h₂) + · rw [cons_inter_of_neg _ h] + exact IH ((le_cons_of_not_mem <| mt (mem_of_le h₂) h).1 h₁) h₂ + +@[simp] +lemma mem_inter : a ∈ s ∩ t ↔ a ∈ s ∧ a ∈ t := + ⟨fun h => ⟨mem_of_le inter_le_left h, mem_of_le inter_le_right h⟩, fun ⟨h₁, h₂⟩ => by + rw [← cons_erase h₁, cons_inter_of_pos _ h₂]; apply mem_cons_self⟩ + +instance instLattice : Lattice (Multiset α) where + sup := (· ∪ ·) + sup_le _ _ _ := union_le + le_sup_left _ _ := le_union_left + le_sup_right _ _ := le_union_right + inf := (· ∩ ·) + le_inf _ _ _ := le_inter + inf_le_left _ _ := inter_le_left + inf_le_right _ _ := inter_le_right + +@[simp] lemma sup_eq_union (s t : Multiset α) : s ⊔ t = s ∪ t := rfl +@[simp] lemma inf_eq_inter (s t : Multiset α) : s ⊓ t = s ∩ t := rfl + +@[simp] lemma le_inter_iff : s ≤ t ∩ u ↔ s ≤ t ∧ s ≤ u := le_inf_iff +@[simp] lemma union_le_iff : s ∪ t ≤ u ↔ s ≤ u ∧ t ≤ u := sup_le_iff + +lemma union_comm (s t : Multiset α) : s ∪ t = t ∪ s := sup_comm .. +lemma inter_comm (s t : Multiset α) : s ∩ t = t ∩ s := inf_comm .. + +lemma eq_union_right (h : s ≤ t) : s ∪ t = t := by rw [union_comm, eq_union_left h] + +@[gcongr] lemma union_le_union_left (h : s ≤ t) (u) : u ∪ s ≤ u ∪ t := sup_le_sup_left h _ + +lemma union_le_add (s t : Multiset α) : s ∪ t ≤ s + t := union_le (le_add_right ..) (le_add_left ..) + +lemma union_add_distrib (s t u : Multiset α) : s ∪ t + u = s + u ∪ (t + u) := by + simpa [(· ∪ ·), union, eq_comm, Multiset.add_assoc, Multiset.add_left_inj] using + show s + u - (t + u) = s - t by + rw [t.add_comm, Multiset.sub_add_eq_sub_sub, Multiset.add_sub_cancel_right] + +lemma add_union_distrib (s t u : Multiset α) : s + (t ∪ u) = s + t ∪ (s + u) := by + rw [Multiset.add_comm, union_add_distrib, s.add_comm, s.add_comm] + +lemma cons_union_distrib (a : α) (s t : Multiset α) : a ::ₘ (s ∪ t) = a ::ₘ s ∪ a ::ₘ t := by + simpa using add_union_distrib (a ::ₘ 0) s t + +lemma inter_add_distrib (s t u : Multiset α) : s ∩ t + u = (s + u) ∩ (t + u) := by + by_contra! h + obtain ⟨a, ha⟩ := lt_iff_cons_le.1 <| h.lt_of_le <| le_inter + (Multiset.add_le_add_right inter_le_left) (Multiset.add_le_add_right inter_le_right) + rw [← cons_add] at ha + exact (lt_cons_self (s ∩ t) a).not_le <| le_inter + (Multiset.le_of_add_le_add_right (ha.trans inter_le_left)) + (Multiset.le_of_add_le_add_right (ha.trans inter_le_right)) + +lemma add_inter_distrib (s t u : Multiset α) : s + t ∩ u = (s + t) ∩ (s + u) := by + rw [Multiset.add_comm, inter_add_distrib, s.add_comm, s.add_comm] + +lemma cons_inter_distrib (a : α) (s t : Multiset α) : a ::ₘ s ∩ t = (a ::ₘ s) ∩ (a ::ₘ t) := by + simp + +lemma union_add_inter (s t : Multiset α) : s ∪ t + s ∩ t = s + t := by + apply _root_.le_antisymm + · rw [union_add_distrib] + refine union_le (Multiset.add_le_add_left inter_le_right) ?_ + rw [Multiset.add_comm] + exact Multiset.add_le_add_right inter_le_left + · rw [Multiset.add_comm, add_inter_distrib] + refine le_inter (Multiset.add_le_add_right le_union_right) ?_ + rw [Multiset.add_comm] + exact Multiset.add_le_add_right le_union_left + +lemma sub_add_inter (s t : Multiset α) : s - t + s ∩ t = s := by + rw [inter_comm] + revert s; refine Multiset.induction_on t (by simp) fun a t IH s => ?_ + by_cases h : a ∈ s + · rw [cons_inter_of_pos _ h, sub_cons, add_cons, IH, cons_erase h] + · rw [cons_inter_of_neg _ h, sub_cons, erase_of_not_mem h, IH] + +lemma sub_inter (s t : Multiset α) : s - s ∩ t = s - t := + (Multiset.eq_sub_of_add_eq <| sub_add_inter ..).symm + +@[simp] +lemma count_inter (a : α) (s t : Multiset α) : count a (s ∩ t) = min (count a s) (count a t) := by + apply @Nat.add_left_cancel (count a (s - t)) + rw [← count_add, sub_add_inter, count_sub, Nat.sub_add_min_cancel] + +@[simp] +lemma coe_inter (s t : List α) : (s ∩ t : Multiset α) = (s.bagInter t : List α) := by ext; simp + +instance instDistribLattice : DistribLattice (Multiset α) where + le_sup_inf s t u := ge_of_eq <| ext.2 fun a ↦ by + simp only [max_min_distrib_left, Multiset.count_inter, Multiset.sup_eq_union, + Multiset.count_union, Multiset.inf_eq_inter] + +@[simp] lemma filter_inter (p : α → Prop) [DecidablePred p] (s t : Multiset α) : + filter p (s ∩ t) = filter p s ∩ filter p t := + le_antisymm (le_inter (filter_le_filter _ inter_le_left) (filter_le_filter _ inter_le_right)) <| + le_filter.2 ⟨inf_le_inf (filter_le _ _) (filter_le _ _), fun _a h => + of_mem_filter (mem_of_le inter_le_left h)⟩ + @[simp] theorem replicate_inter (n : ℕ) (x : α) (s : Multiset α) : replicate n x ∩ s = replicate (min n (s.count x)) x := by @@ -2330,23 +2231,7 @@ theorem inter_replicate (s : Multiset α) (n : ℕ) (x : α) : s ∩ replicate n x = replicate (min (s.count x) n) x := by rw [inter_comm, replicate_inter, min_comm] -theorem erase_attach_map_val (s : Multiset α) (x : {x // x ∈ s}) : - (s.attach.erase x).map (↑) = s.erase x := by - rw [Multiset.map_erase _ val_injective, attach_map_val] - -theorem erase_attach_map (s : Multiset α) (f : α → β) (x : {x // x ∈ s}) : - (s.attach.erase x).map (fun j : {x // x ∈ s} ↦ f j) = (s.erase x).map f := by - simp only [← Function.comp_apply (f := f)] - rw [← map_map, erase_attach_map_val] - -end - -@[ext] -theorem addHom_ext [AddZeroClass β] ⦃f g : Multiset α →+ β⦄ (h : ∀ x, f {x} = g {x}) : f = g := by - ext s - induction' s using Multiset.induction_on with a s ih - · simp only [_root_.map_zero] - · simp only [← singleton_add, _root_.map_add, ih, h] +end sub section Embedding @@ -2702,7 +2587,7 @@ lemma add_eq_union_left_of_le [DecidableEq α] {s t u : Multiset α} (h : t ≤ rw [← add_eq_union_iff_disjoint] refine ⟨fun h0 ↦ ?_, ?_⟩ · rw [and_iff_right_of_imp] - · exact (le_of_add_le_add_left <| h0.trans_le <| union_le_add u t).antisymm h + · exact (Multiset.le_of_add_le_add_left <| h0.trans_le <| union_le_add u t).antisymm h · rintro rfl exact h0 · rintro ⟨h0, rfl⟩ @@ -2800,4 +2685,4 @@ theorem coe_subsingletonEquiv [Subsingleton α] : end Multiset -set_option linter.style.longFile 2900 +set_option linter.style.longFile 2800 diff --git a/Mathlib/Data/Multiset/Bind.lean b/Mathlib/Data/Multiset/Bind.lean index b34913bef0ace..766b7f70ea703 100644 --- a/Mathlib/Data/Multiset/Bind.lean +++ b/Mathlib/Data/Multiset/Bind.lean @@ -201,6 +201,7 @@ theorem attach_bind_coe (s : Multiset α) (f : α → Multiset β) : variable {f s t} +open scoped Function in -- required for scoped `on` notation @[simp] lemma nodup_bind : Nodup (bind s f) ↔ (∀ a ∈ s, Nodup (f a)) ∧ s.Pairwise (Disjoint on f) := by have : ∀ a, ∃ l : List β, f a = l := fun a => Quot.induction_on (f a) fun l => ⟨l, rfl⟩ diff --git a/Mathlib/Data/Multiset/Dedup.lean b/Mathlib/Data/Multiset/Dedup.lean index 7e580eb8e459b..07c9eabf1192c 100644 --- a/Mathlib/Data/Multiset/Dedup.lean +++ b/Mathlib/Data/Multiset/Dedup.lean @@ -10,6 +10,7 @@ import Mathlib.Data.Multiset.Nodup # Erasing duplicates in a multiset. -/ +assert_not_exists Monoid namespace Multiset @@ -104,11 +105,6 @@ theorem dedup_map_dedup_eq [DecidableEq β] (f : α → β) (s : Multiset α) : dedup (map f (dedup s)) = dedup (map f s) := by simp [dedup_ext] -@[simp] -theorem dedup_nsmul {s : Multiset α} {n : ℕ} (h0 : n ≠ 0) : (n • s).dedup = s.dedup := by - ext a - by_cases h : a ∈ s <;> simp [h, h0] - theorem Nodup.le_dedup_iff_le {s t : Multiset α} (hno : s.Nodup) : s ≤ t.dedup ↔ s ≤ t := by simp [le_dedup, hno] @@ -119,7 +115,7 @@ theorem Subset.dedup_add_right {s t : Multiset α} (h : s ⊆ t) : theorem Subset.dedup_add_left {s t : Multiset α} (h : t ⊆ s) : dedup (s + t) = dedup s := by - rw [add_comm, Subset.dedup_add_right h] + rw [s.add_comm, Subset.dedup_add_right h] theorem Disjoint.dedup_add {s t : Multiset α} (h : Disjoint s t) : dedup (s + t) = dedup s + dedup t := by @@ -132,9 +128,3 @@ theorem _root_.List.Subset.dedup_append_left {s t : List α} (h : t ⊆ s) : rw [← coe_eq_coe, ← coe_dedup, ← coe_add, Subset.dedup_add_left h, coe_dedup] end Multiset - -theorem Multiset.Nodup.le_nsmul_iff_le {α : Type*} {s t : Multiset α} {n : ℕ} (h : s.Nodup) - (hn : n ≠ 0) : s ≤ n • t ↔ s ≤ t := by - classical - rw [← h.le_dedup_iff_le, Iff.comm, ← h.le_dedup_iff_le] - simp [hn] diff --git a/Mathlib/Data/Multiset/NatAntidiagonal.lean b/Mathlib/Data/Multiset/NatAntidiagonal.lean index b9d75740e9a1f..55e2353bf8752 100644 --- a/Mathlib/Data/Multiset/NatAntidiagonal.lean +++ b/Mathlib/Data/Multiset/NatAntidiagonal.lean @@ -19,6 +19,8 @@ This refines file `Data.List.NatAntidiagonal` and is further refined by file `Data.Finset.NatAntidiagonal`. -/ +assert_not_exists Monoid + namespace Multiset namespace Nat @@ -55,8 +57,8 @@ theorem antidiagonal_succ {n : ℕ} : theorem antidiagonal_succ' {n : ℕ} : antidiagonal (n + 1) = (n + 1, 0) ::ₘ (antidiagonal n).map (Prod.map id Nat.succ) := by - rw [antidiagonal, List.Nat.antidiagonal_succ', ← coe_add, add_comm, antidiagonal, map_coe, - coe_add, List.singleton_append, cons_coe] + rw [antidiagonal, List.Nat.antidiagonal_succ', ← coe_add, Multiset.add_comm, antidiagonal, + map_coe, coe_add, List.singleton_append, cons_coe] theorem antidiagonal_succ_succ' {n : ℕ} : antidiagonal (n + 2) = diff --git a/Mathlib/Data/Multiset/Nodup.lean b/Mathlib/Data/Multiset/Nodup.lean index 87bbfbd63a93a..66e5cd4680dd3 100644 --- a/Mathlib/Data/Multiset/Nodup.lean +++ b/Mathlib/Data/Multiset/Nodup.lean @@ -10,6 +10,7 @@ import Mathlib.Data.List.Pairwise # The `Nodup` predicate for multisets without duplicate elements. -/ +assert_not_exists Monoid namespace Multiset @@ -62,7 +63,7 @@ theorem nodup_iff_ne_cons_cons {s : Multiset α} : s.Nodup ↔ ∀ a t, s ≠ a nodup_iff_le.trans ⟨fun h a _ s_eq => h a (s_eq.symm ▸ cons_le_cons a (cons_le_cons a (zero_le _))), fun h a le => let ⟨t, s_eq⟩ := le_iff_exists_add.mp le - h a t (by rwa [cons_add, cons_add, zero_add] at s_eq)⟩ + h a t (by rwa [cons_add, cons_add, Multiset.zero_add] at s_eq)⟩ theorem nodup_iff_count_le_one [DecidableEq α] {s : Multiset α} : Nodup s ↔ ∀ a, count a s ≤ 1 := Quot.induction_on s fun _l => by @@ -193,7 +194,7 @@ theorem mem_sub_of_nodup [DecidableEq α] {a : α} {s t : Multiset α} (d : Nodu refine count_eq_zero.1 ?_ h rw [count_sub a s t, Nat.sub_eq_zero_iff_le] exact le_trans (nodup_iff_count_le_one.1 d _) (count_pos.2 h')⟩, - fun ⟨h₁, h₂⟩ => Or.resolve_right (mem_add.1 <| mem_of_le le_tsub_add h₁) h₂⟩ + fun ⟨h₁, h₂⟩ => Or.resolve_right (mem_add.1 <| mem_of_le Multiset.le_sub_add h₁) h₂⟩ theorem map_eq_map_of_bij_of_nodup (f : α → γ) (g : β → γ) {s : Multiset α} {t : Multiset β} (hs : s.Nodup) (ht : t.Nodup) (i : ∀ a ∈ s, β) (hi : ∀ a ha, i a ha ∈ t) diff --git a/Mathlib/Data/Multiset/OrderedMonoid.lean b/Mathlib/Data/Multiset/OrderedMonoid.lean index bf40bc0b213eb..a674512d9f22c 100644 --- a/Mathlib/Data/Multiset/OrderedMonoid.lean +++ b/Mathlib/Data/Multiset/OrderedMonoid.lean @@ -3,7 +3,7 @@ Copyright (c) 2015 Microsoft Corporation. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ -import Mathlib.Data.Multiset.Basic +import Mathlib.Algebra.Order.Group.Multiset import Mathlib.Algebra.Order.Monoid.Canonical.Defs /-! diff --git a/Mathlib/Data/Multiset/Range.lean b/Mathlib/Data/Multiset/Range.lean index 7f1b911b0f826..61399e9389712 100644 --- a/Mathlib/Data/Multiset/Range.lean +++ b/Mathlib/Data/Multiset/Range.lean @@ -7,6 +7,7 @@ import Mathlib.Data.Multiset.Basic /-! # `Multiset.range n` gives `{0, 1, ..., n-1}` as a multiset. -/ +assert_not_exists Monoid open List Nat @@ -27,7 +28,7 @@ theorem range_zero : range 0 = 0 := @[simp] theorem range_succ (n : ℕ) : range (succ n) = n ::ₘ range n := by - rw [range, List.range_succ, ← coe_add, add_comm]; rfl + rw [range, List.range_succ, ← coe_add, Multiset.add_comm]; rfl @[simp] theorem card_range (n : ℕ) : card (range n) = n := diff --git a/Mathlib/Data/Multiset/Sum.lean b/Mathlib/Data/Multiset/Sum.lean index 0dd8540893604..66512fa587909 100644 --- a/Mathlib/Data/Multiset/Sum.lean +++ b/Mathlib/Data/Multiset/Sum.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ +import Mathlib.Algebra.Order.Group.Multiset import Mathlib.Data.Multiset.Nodup /-! @@ -29,11 +30,11 @@ def disjSum : Multiset (α ⊕ β) := @[simp] theorem zero_disjSum : (0 : Multiset α).disjSum t = t.map inr := - zero_add _ + Multiset.zero_add _ @[simp] theorem disjSum_zero : s.disjSum (0 : Multiset β) = s.map inl := - add_zero _ + Multiset.add_zero _ @[simp] theorem card_disjSum : Multiset.card (s.disjSum t) = Multiset.card s + Multiset.card t := by @@ -64,11 +65,11 @@ theorem disjSum_mono (hs : s₁ ≤ s₂) (ht : t₁ ≤ t₂) : s₁.disjSum t add_le_add (map_le_map hs) (map_le_map ht) theorem disjSum_mono_left (t : Multiset β) : Monotone fun s : Multiset α => s.disjSum t := - fun _ _ hs => add_le_add_right (map_le_map hs) _ + fun _ _ hs => Multiset.add_le_add_right (map_le_map hs) theorem disjSum_mono_right (s : Multiset α) : Monotone (s.disjSum : Multiset β → Multiset (α ⊕ β)) := fun _ _ ht => - add_le_add_left (map_le_map ht) _ + Multiset.add_le_add_left (map_le_map ht) theorem disjSum_lt_disjSum_of_lt_of_le (hs : s₁ < s₂) (ht : t₁ ≤ t₂) : s₁.disjSum t₁ < s₂.disjSum t₂ := diff --git a/Mathlib/Data/NNRat/Defs.lean b/Mathlib/Data/NNRat/Defs.lean index 1b559909a9913..05d7636202277 100644 --- a/Mathlib/Data/NNRat/Defs.lean +++ b/Mathlib/Data/NNRat/Defs.lean @@ -188,7 +188,6 @@ def coeHom : ℚ≥0 →+* ℚ where @[simp, norm_cast] lemma coe_natCast (n : ℕ) : (↑(↑n : ℚ≥0) : ℚ) = n := rfl --- See note [no_index around OfNat.ofNat] @[simp] theorem mk_natCast (n : ℕ) : @Eq ℚ≥0 (⟨(n : ℚ), Nat.cast_nonneg' n⟩ : ℚ≥0) n := rfl @@ -336,10 +335,9 @@ lemma coprime_num_den (q : ℚ≥0) : q.num.Coprime q.den := by simpa [num, den] @[simp, norm_cast] lemma num_natCast (n : ℕ) : num n = n := rfl @[simp, norm_cast] lemma den_natCast (n : ℕ) : den n = 1 := rfl --- See note [no_index around OfNat.ofNat] -@[simp] lemma num_ofNat (n : ℕ) [n.AtLeastTwo] : num (no_index (OfNat.ofNat n)) = OfNat.ofNat n := +@[simp] lemma num_ofNat (n : ℕ) [n.AtLeastTwo] : num ofNat(n) = OfNat.ofNat n := rfl -@[simp] lemma den_ofNat (n : ℕ) [n.AtLeastTwo] : den (no_index (OfNat.ofNat n)) = 1 := rfl +@[simp] lemma den_ofNat (n : ℕ) [n.AtLeastTwo] : den ofNat(n) = 1 := rfl theorem ext_num_den (hn : p.num = q.num) (hd : p.den = q.den) : p = q := by refine ext <| Rat.ext ?_ hd diff --git a/Mathlib/Data/NNReal/Defs.lean b/Mathlib/Data/NNReal/Defs.lean index c773102bcabad..dbf988eb67926 100644 --- a/Mathlib/Data/NNReal/Defs.lean +++ b/Mathlib/Data/NNReal/Defs.lean @@ -143,7 +143,6 @@ protected theorem coe_injective : Injective ((↑) : ℝ≥0 → ℝ) := Subtype @[simp, norm_cast] lemma coe_inj {r₁ r₂ : ℝ≥0} : (r₁ : ℝ) = r₂ ↔ r₁ = r₂ := NNReal.coe_injective.eq_iff -@[deprecated (since := "2024-02-03")] protected alias coe_eq := coe_inj @[simp, norm_cast] lemma coe_zero : ((0 : ℝ≥0) : ℝ) = 0 := rfl @@ -208,13 +207,13 @@ theorem smul_def {M : Type*} [MulAction ℝ M] (c : ℝ≥0) (x : M) : c • x = rfl instance {M N : Type*} [MulAction ℝ M] [MulAction ℝ N] [SMul M N] [IsScalarTower ℝ M N] : - IsScalarTower ℝ≥0 M N where smul_assoc r := (smul_assoc (r : ℝ) : _) + IsScalarTower ℝ≥0 M N where smul_assoc r := smul_assoc (r : ℝ) instance smulCommClass_left {M N : Type*} [MulAction ℝ N] [SMul M N] [SMulCommClass ℝ M N] : - SMulCommClass ℝ≥0 M N where smul_comm r := (smul_comm (r : ℝ) : _) + SMulCommClass ℝ≥0 M N where smul_comm r := smul_comm (r : ℝ) instance smulCommClass_right {M N : Type*} [MulAction ℝ N] [SMul M N] [SMulCommClass M ℝ N] : - SMulCommClass M ℝ≥0 N where smul_comm m r := (smul_comm m (r : ℝ) : _) + SMulCommClass M ℝ≥0 N where smul_comm m r := smul_comm m (r : ℝ) /-- A `DistribMulAction` over `ℝ` restricts to a `DistribMulAction` over `ℝ≥0`. -/ instance {M : Type*} [AddMonoid M] [DistribMulAction ℝ M] : DistribMulAction ℝ≥0 M := @@ -315,10 +314,9 @@ theorem mk_natCast (n : ℕ) : @Eq ℝ≥0 (⟨(n : ℝ), n.cast_nonneg⟩ : ℝ theorem toNNReal_coe_nat (n : ℕ) : Real.toNNReal n = n := NNReal.eq <| by simp [Real.coe_toNNReal] --- See note [no_index around OfNat.ofNat] @[simp] theorem _root_.Real.toNNReal_ofNat (n : ℕ) [n.AtLeastTwo] : - Real.toNNReal (no_index (OfNat.ofNat n)) = OfNat.ofNat n := + Real.toNNReal ofNat(n) = OfNat.ofNat n := toNNReal_coe_nat n /-- `Real.toNNReal` and `NNReal.toReal : ℝ≥0 → ℝ` form a Galois insertion. -/ @@ -480,7 +478,7 @@ theorem zero_le_coe {q : ℝ≥0} : 0 ≤ (q : ℝ) := instance instOrderedSMul {M : Type*} [OrderedAddCommMonoid M] [Module ℝ M] [OrderedSMul ℝ M] : OrderedSMul ℝ≥0 M where - smul_lt_smul_of_pos hab hc := (smul_lt_smul_of_pos_left hab (NNReal.coe_pos.2 hc) : _) + smul_lt_smul_of_pos hab hc := (smul_lt_smul_of_pos_left hab (NNReal.coe_pos.2 hc) :) lt_of_smul_lt_smul_of_pos {_ _ c} hab _ := lt_of_smul_lt_smul_of_nonneg_left (by exact hab) (NNReal.coe_nonneg c) @@ -529,7 +527,7 @@ alias toNNReal_eq_nat_cast := toNNReal_eq_natCast @[simp] lemma toNNReal_eq_ofNat {r : ℝ} {n : ℕ} [n.AtLeastTwo] : - r.toNNReal = no_index (OfNat.ofNat n) ↔ r = OfNat.ofNat n := + r.toNNReal = ofNat(n) ↔ r = OfNat.ofNat n := toNNReal_eq_natCast (NeZero.ne n) @[simp] @@ -560,12 +558,12 @@ alias nat_cast_lt_toNNReal := natCast_lt_toNNReal @[simp] lemma toNNReal_le_ofNat {r : ℝ} {n : ℕ} [n.AtLeastTwo] : - r.toNNReal ≤ no_index (OfNat.ofNat n) ↔ r ≤ n := + r.toNNReal ≤ ofNat(n) ↔ r ≤ n := toNNReal_le_natCast @[simp] lemma ofNat_lt_toNNReal {r : ℝ} {n : ℕ} [n.AtLeastTwo] : - no_index (OfNat.ofNat n) < r.toNNReal ↔ n < r := + ofNat(n) < r.toNNReal ↔ n < r := natCast_lt_toNNReal @[simp] @@ -626,12 +624,12 @@ alias toNNReal_lt_nat_cast := toNNReal_lt_natCast @[simp] lemma toNNReal_lt_ofNat {r : ℝ} {n : ℕ} [n.AtLeastTwo] : - r.toNNReal < no_index (OfNat.ofNat n) ↔ r < OfNat.ofNat n := + r.toNNReal < ofNat(n) ↔ r < OfNat.ofNat n := toNNReal_lt_natCast (NeZero.ne n) @[simp] lemma ofNat_le_toNNReal {n : ℕ} {r : ℝ} [n.AtLeastTwo] : - no_index (OfNat.ofNat n) ≤ r.toNNReal ↔ OfNat.ofNat n ≤ r := + ofNat(n) ≤ r.toNNReal ↔ OfNat.ofNat n ≤ r := natCast_le_toNNReal (NeZero.ne n) @[simp] diff --git a/Mathlib/Data/Nat/Cast/Basic.lean b/Mathlib/Data/Nat/Cast/Basic.lean index 77459f473b112..4d4457e98263a 100644 --- a/Mathlib/Data/Nat/Cast/Basic.lean +++ b/Mathlib/Data/Nat/Cast/Basic.lean @@ -158,12 +158,12 @@ theorem map_natCast [FunLike F R S] [RingHomClass F R S] (f : F) : ∀ n : ℕ, map_natCast' f <| map_one f /-- This lemma is not marked `@[simp]` lemma because its `#discr_tree_key` (for the LHS) would just -be `DFunLike.coe _ _`, due to the `no_index` that https://github.com/leanprover/lean4/issues/2867 +be `DFunLike.coe _ _`, due to the `ofNat` that https://github.com/leanprover/lean4/issues/2867 forces us to include, and therefore it would negatively impact performance. If that issue is resolved, this can be marked `@[simp]`. -/ theorem map_ofNat [FunLike F R S] [RingHomClass F R S] (f : F) (n : ℕ) [Nat.AtLeastTwo n] : - (f (no_index (OfNat.ofNat n)) : S) = OfNat.ofNat n := + (f ofNat(n) : S) = OfNat.ofNat n := map_natCast f n theorem ext_nat [FunLike F ℕ R] [RingHomClass F ℕ R] (f g : F) : f = g := diff --git a/Mathlib/Data/Nat/Cast/Order/Basic.lean b/Mathlib/Data/Nat/Cast/Order/Basic.lean index eec03f2914acb..6c96978eb5a1b 100644 --- a/Mathlib/Data/Nat/Cast/Order/Basic.lean +++ b/Mathlib/Data/Nat/Cast/Order/Basic.lean @@ -33,9 +33,6 @@ theorem mono_cast : Monotone (Nat.cast : ℕ → α) := monotone_nat_of_le_succ fun n ↦ by rw [Nat.cast_succ]; exact le_add_of_nonneg_right zero_le_one -@[deprecated mono_cast (since := "2024-02-10")] -theorem cast_le_cast {a b : ℕ} (h : a ≤ b) : (a : α) ≤ b := mono_cast h - @[gcongr] theorem _root_.GCongr.natCast_le_natCast {a b : ℕ} (h : a ≤ b) : (a : α) ≤ b := mono_cast h @@ -45,9 +42,8 @@ theorem cast_nonneg' (n : ℕ) : 0 ≤ (n : α) := @Nat.cast_zero α _ ▸ mono_cast (Nat.zero_le n) /-- See also `Nat.ofNat_nonneg`, specialised for an `OrderedSemiring`. -/ --- See note [no_index around OfNat.ofNat] @[simp low] -theorem ofNat_nonneg' (n : ℕ) [n.AtLeastTwo] : 0 ≤ (no_index (OfNat.ofNat n : α)) := cast_nonneg' n +theorem ofNat_nonneg' (n : ℕ) [n.AtLeastTwo] : 0 ≤ (ofNat(n) : α) := cast_nonneg' n section Nontrivial @@ -103,48 +99,40 @@ theorem cast_le_one : (n : α) ≤ 1 ↔ n ≤ 1 := by rw [← cast_one, cast_le section variable [m.AtLeastTwo] --- See note [no_index around OfNat.ofNat] @[simp] -theorem ofNat_le_cast : (no_index (OfNat.ofNat m : α)) ≤ n ↔ (OfNat.ofNat m : ℕ) ≤ n := +theorem ofNat_le_cast : (ofNat(m) : α) ≤ n ↔ (OfNat.ofNat m : ℕ) ≤ n := cast_le --- See note [no_index around OfNat.ofNat] @[simp] -theorem ofNat_lt_cast : (no_index (OfNat.ofNat m : α)) < n ↔ (OfNat.ofNat m : ℕ) < n := +theorem ofNat_lt_cast : (ofNat(m) : α) < n ↔ (OfNat.ofNat m : ℕ) < n := cast_lt end variable [n.AtLeastTwo] --- See note [no_index around OfNat.ofNat] @[simp] -theorem cast_le_ofNat : (m : α) ≤ (no_index (OfNat.ofNat n)) ↔ m ≤ OfNat.ofNat n := +theorem cast_le_ofNat : (m : α) ≤ (ofNat(n) : α) ↔ m ≤ OfNat.ofNat n := cast_le --- See note [no_index around OfNat.ofNat] @[simp] -theorem cast_lt_ofNat : (m : α) < (no_index (OfNat.ofNat n)) ↔ m < OfNat.ofNat n := +theorem cast_lt_ofNat : (m : α) < (ofNat(n) : α) ↔ m < OfNat.ofNat n := cast_lt --- See note [no_index around OfNat.ofNat] @[simp] -theorem one_lt_ofNat : 1 < (no_index (OfNat.ofNat n : α)) := +theorem one_lt_ofNat : 1 < (ofNat(n) : α) := one_lt_cast.mpr AtLeastTwo.one_lt --- See note [no_index around OfNat.ofNat] @[simp] -theorem one_le_ofNat : 1 ≤ (no_index (OfNat.ofNat n : α)) := +theorem one_le_ofNat : 1 ≤ (ofNat(n) : α) := one_le_cast.mpr NeZero.one_le --- See note [no_index around OfNat.ofNat] @[simp] -theorem not_ofNat_le_one : ¬(no_index (OfNat.ofNat n : α)) ≤ 1 := +theorem not_ofNat_le_one : ¬(ofNat(n) : α) ≤ 1 := (cast_le_one.not.trans not_le).mpr AtLeastTwo.one_lt --- See note [no_index around OfNat.ofNat] @[simp] -theorem not_ofNat_lt_one : ¬(no_index (OfNat.ofNat n : α)) < 1 := +theorem not_ofNat_lt_one : ¬(ofNat(n) : α) < 1 := mt le_of_lt not_ofNat_le_one variable [m.AtLeastTwo] @@ -153,18 +141,14 @@ variable [m.AtLeastTwo] -- and `Nat.cast_ofNat`, but their LHSs match literally every inequality, so they're too expensive. -- If https://github.com/leanprover/lean4/issues/2867 is fixed in a performant way, these can be made `@[simp]`. --- See note [no_index around OfNat.ofNat] -- @[simp] theorem ofNat_le : - (no_index (OfNat.ofNat m : α)) ≤ (no_index (OfNat.ofNat n)) ↔ - (OfNat.ofNat m : ℕ) ≤ OfNat.ofNat n := + (ofNat(m) : α) ≤ (ofNat(n) : α) ↔ (OfNat.ofNat m : ℕ) ≤ OfNat.ofNat n := cast_le --- See note [no_index around OfNat.ofNat] -- @[simp] theorem ofNat_lt : - (no_index (OfNat.ofNat m : α)) < (no_index (OfNat.ofNat n)) ↔ - (OfNat.ofNat m : ℕ) < OfNat.ofNat n := + (ofNat(m) : α) < (ofNat(n) : α) ↔ (OfNat.ofNat m : ℕ) < OfNat.ofNat n := cast_lt end OrderedSemiring diff --git a/Mathlib/Data/Nat/Cast/Synonym.lean b/Mathlib/Data/Nat/Cast/Synonym.lean index 11b1f4d3b233a..c18a4c25a81dc 100644 --- a/Mathlib/Data/Nat/Cast/Synonym.lean +++ b/Mathlib/Data/Nat/Cast/Synonym.lean @@ -74,7 +74,7 @@ theorem toLex_natCast [NatCast α] (n : ℕ) : toLex (n : α) = n := @[simp] theorem toLex_ofNat [NatCast α] (n : ℕ) [n.AtLeastTwo] : - (toLex (no_index (OfNat.ofNat n : α))) = OfNat.ofNat n := + toLex (ofNat(n) : α) = OfNat.ofNat n := rfl @[simp] @@ -83,5 +83,5 @@ theorem ofLex_natCast [NatCast α] (n : ℕ) : (ofLex n : α) = n := @[simp] theorem ofLex_ofNat [NatCast α] (n : ℕ) [n.AtLeastTwo] : - (ofLex (no_index (OfNat.ofNat n : Lex α))) = OfNat.ofNat n := + ofLex (ofNat(n) : Lex α) = OfNat.ofNat n := rfl diff --git a/Mathlib/Data/Nat/ChineseRemainder.lean b/Mathlib/Data/Nat/ChineseRemainder.lean index 17beeef782677..e7147fe88bf42 100644 --- a/Mathlib/Data/Nat/ChineseRemainder.lean +++ b/Mathlib/Data/Nat/ChineseRemainder.lean @@ -22,6 +22,7 @@ Gödel's Beta function, which is used in proving Gödel's incompleteness theorem Chinese Remainder Theorem, Gödel, beta function -/ +open scoped Function -- required for scoped `on` notation namespace Nat variable {ι : Type*} diff --git a/Mathlib/Data/Nat/Choose/Multinomial.lean b/Mathlib/Data/Nat/Choose/Multinomial.lean index 318cee9f748cd..ea155b24f4e44 100644 --- a/Mathlib/Data/Nat/Choose/Multinomial.lean +++ b/Mathlib/Data/Nat/Choose/Multinomial.lean @@ -207,6 +207,8 @@ variable {α R : Type*} [DecidableEq α] section Semiring variable [Semiring R] +open scoped Function -- required for scoped `on` notation + -- TODO: Can we prove one of the following two from the other one? /-- The **multinomial theorem**. -/ lemma sum_pow_eq_sum_piAntidiag_of_commute (s : Finset α) (f : α → R) diff --git a/Mathlib/Data/Nat/Digits.lean b/Mathlib/Data/Nat/Digits.lean index 5adc574882628..bde0bf0b51b40 100644 --- a/Mathlib/Data/Nat/Digits.lean +++ b/Mathlib/Data/Nat/Digits.lean @@ -740,7 +740,6 @@ lemma toDigitsCore_lens_eq_aux (b f : Nat) : specialize ih (n / b) (Nat.digitChar (n % b) :: l1) (Nat.digitChar (n % b) :: l2) simp only [List.length, congrArg (fun l ↦ l + 1) hlen] at ih exact ih trivial -@[deprecated (since := "2024-02-19")] alias to_digits_core_lens_eq_aux:= toDigitsCore_lens_eq_aux lemma toDigitsCore_lens_eq (b f : Nat) : ∀ (n : Nat) (c : Char) (tl : List Char), (Nat.toDigitsCore b f n (c :: tl)).length = (Nat.toDigitsCore b f n tl).length + 1 := by @@ -757,7 +756,6 @@ lemma toDigitsCore_lens_eq (b f : Nat) : ∀ (n : Nat) (c : Char) (tl : List Cha have lens_eq : (x :: (c :: tl)).length = (c :: x :: tl).length := by simp apply toDigitsCore_lens_eq_aux exact lens_eq -@[deprecated (since := "2024-02-19")] alias to_digits_core_lens_eq:= toDigitsCore_lens_eq lemma nat_repr_len_aux (n b e : Nat) (h_b_pos : 0 < b) : n < b ^ e.succ → n / b < b ^ e := by simp only [Nat.pow_succ] @@ -788,7 +786,6 @@ lemma toDigitsCore_length (b : Nat) (h : 2 <= b) (f n e : Nat) have _ : b ^ 1 = b := by simp only [Nat.pow_succ, pow_zero, Nat.one_mul] have _ : n < b := ‹b ^ 1 = b› ▸ hlt simp [(@Nat.div_eq_of_lt n b ‹n < b› : n / b = 0)] -@[deprecated (since := "2024-02-19")] alias to_digits_core_length := toDigitsCore_length /-- The core implementation of `Nat.repr` returns a String with length less than or equal to the number of digits in the decimal number (represented by `e`). For example, the decimal string diff --git a/Mathlib/Data/Nat/PartENat.lean b/Mathlib/Data/Nat/PartENat.lean index 1000c54d34db3..22dbf8f8562ca 100644 --- a/Mathlib/Data/Nat/PartENat.lean +++ b/Mathlib/Data/Nat/PartENat.lean @@ -114,9 +114,8 @@ theorem natCast_inj {x y : ℕ} : (x : PartENat) = y ↔ x = y := theorem dom_natCast (x : ℕ) : (x : PartENat).Dom := trivial --- See note [no_index around OfNat.ofNat] @[simp] -theorem dom_ofNat (x : ℕ) [x.AtLeastTwo] : (no_index (OfNat.ofNat x : PartENat)).Dom := +theorem dom_ofNat (x : ℕ) [x.AtLeastTwo] : (ofNat(x) : PartENat).Dom := trivial @[simp] @@ -189,10 +188,9 @@ theorem get_zero (h : (0 : PartENat).Dom) : (0 : PartENat).get h = 0 := theorem get_one (h : (1 : PartENat).Dom) : (1 : PartENat).get h = 1 := rfl --- See note [no_index around OfNat.ofNat] @[simp] -theorem get_ofNat' (x : ℕ) [x.AtLeastTwo] (h : (no_index (OfNat.ofNat x : PartENat)).Dom) : - Part.get (no_index (OfNat.ofNat x : PartENat)) h = (no_index (OfNat.ofNat x)) := +theorem get_ofNat' (x : ℕ) [x.AtLeastTwo] (h : (ofNat(x) : PartENat).Dom) : + Part.get (ofNat(x) : PartENat) h = ofNat(x) := get_natCast' x h nonrec theorem get_eq_iff_eq_some {a : PartENat} {ha : a.Dom} {b : ℕ} : a.get ha = b ↔ a = some b := @@ -327,9 +325,8 @@ theorem zero_lt_top : (0 : PartENat) < ⊤ := theorem one_lt_top : (1 : PartENat) < ⊤ := natCast_lt_top 1 --- See note [no_index around OfNat.ofNat] @[simp] -theorem ofNat_lt_top (x : ℕ) [x.AtLeastTwo] : (no_index (OfNat.ofNat x : PartENat)) < ⊤ := +theorem ofNat_lt_top (x : ℕ) [x.AtLeastTwo] : (ofNat(x) : PartENat) < ⊤ := natCast_lt_top x @[simp] @@ -344,9 +341,8 @@ theorem zero_ne_top : (0 : PartENat) ≠ ⊤ := theorem one_ne_top : (1 : PartENat) ≠ ⊤ := natCast_ne_top 1 --- See note [no_index around OfNat.ofNat] @[simp] -theorem ofNat_ne_top (x : ℕ) [x.AtLeastTwo] : (no_index (OfNat.ofNat x : PartENat)) ≠ ⊤ := +theorem ofNat_ne_top (x : ℕ) [x.AtLeastTwo] : (ofNat(x) : PartENat) ≠ ⊤ := natCast_ne_top x theorem not_isMax_natCast (x : ℕ) : ¬IsMax (x : PartENat) := @@ -556,7 +552,7 @@ theorem toWithTop_natCast' (n : ℕ) {_ : Decidable (n : PartENat).Dom} : @[simp] theorem toWithTop_ofNat (n : ℕ) [n.AtLeastTwo] {_ : Decidable (OfNat.ofNat n : PartENat).Dom} : - toWithTop (no_index (OfNat.ofNat n : PartENat)) = OfNat.ofNat n := toWithTop_natCast' n + toWithTop (ofNat(n) : PartENat) = OfNat.ofNat n := toWithTop_natCast' n -- Porting note: statement changed. Mathlib 3 statement was -- ``` @@ -621,7 +617,7 @@ theorem ofENat_zero : ofENat 0 = 0 := rfl theorem ofENat_one : ofENat 1 = 1 := rfl @[simp, norm_cast] -theorem ofENat_ofNat (n : Nat) [n.AtLeastTwo] : ofENat (no_index (OfNat.ofNat n)) = OfNat.ofNat n := +theorem ofENat_ofNat (n : Nat) [n.AtLeastTwo] : ofENat ofNat(n) = OfNat.ofNat n := rfl @[simp, norm_cast] @@ -679,7 +675,7 @@ theorem withTopEquiv_one : withTopEquiv 1 = 1 := by simp theorem withTopEquiv_ofNat (n : Nat) [n.AtLeastTwo] : - withTopEquiv (no_index (OfNat.ofNat n)) = OfNat.ofNat n := by + withTopEquiv ofNat(n) = OfNat.ofNat n := by simp theorem withTopEquiv_le {x y : PartENat} : withTopEquiv x ≤ withTopEquiv y ↔ x ≤ y := by @@ -701,7 +697,7 @@ theorem withTopEquiv_symm_one : withTopEquiv.symm 1 = 1 := by simp theorem withTopEquiv_symm_ofNat (n : Nat) [n.AtLeastTwo] : - withTopEquiv.symm (no_index (OfNat.ofNat n)) = OfNat.ofNat n := by + withTopEquiv.symm ofNat(n) = OfNat.ofNat n := by simp theorem withTopEquiv_symm_le {x y : ℕ∞} : withTopEquiv.symm x ≤ withTopEquiv.symm y ↔ x ≤ y := by diff --git a/Mathlib/Data/Nat/Prime/Int.lean b/Mathlib/Data/Nat/Prime/Int.lean index d357a3d303050..474b170148988 100644 --- a/Mathlib/Data/Nat/Prime/Int.lean +++ b/Mathlib/Data/Nat/Prime/Int.lean @@ -38,10 +38,9 @@ end Nat namespace Int --- See note [no_index around OfNat.ofNat] @[simp] theorem prime_ofNat_iff {n : ℕ} : - Prime (no_index (OfNat.ofNat n : ℤ)) ↔ Nat.Prime (OfNat.ofNat n) := + Prime (ofNat(n) : ℤ) ↔ Nat.Prime (OfNat.ofNat n) := Nat.prime_iff_prime_int.symm theorem prime_two : Prime (2 : ℤ) := diff --git a/Mathlib/Data/PFunctor/Multivariate/Basic.lean b/Mathlib/Data/PFunctor/Multivariate/Basic.lean index 319632029d761..30e5c5bdb6316 100644 --- a/Mathlib/Data/PFunctor/Multivariate/Basic.lean +++ b/Mathlib/Data/PFunctor/Multivariate/Basic.lean @@ -109,7 +109,7 @@ end Const /-- Functor composition on polynomial functors -/ def comp (P : MvPFunctor.{u} n) (Q : Fin2 n → MvPFunctor.{u} m) : MvPFunctor m where A := Σ a₂ : P.1, ∀ i, P.2 a₂ i → (Q i).1 - B a i := Σ(j : _) (b : P.2 a.1 j), (Q j).2 (a.snd j b) i + B a i := Σ (j : _) (b : P.2 a.1 j), (Q j).2 (a.snd j b) i variable {P} {Q : Fin2 n → MvPFunctor.{u} m} {α β : TypeVec.{u} m} diff --git a/Mathlib/Data/PFunctor/Multivariate/M.lean b/Mathlib/Data/PFunctor/Multivariate/M.lean index cb6ed15fd11ac..758d698261463 100644 --- a/Mathlib/Data/PFunctor/Multivariate/M.lean +++ b/Mathlib/Data/PFunctor/Multivariate/M.lean @@ -238,7 +238,7 @@ theorem M.bisim {α : TypeVec n} (R : P.M α → P.M α → Prop) rcases M.bisim_lemma P e₂ with ⟨g₂', e₂', e₃, rfl⟩ cases h'.symm.trans e₁' cases h'.symm.trans e₂' - · exact (congr_fun (congr_fun e₃ i) c : _) + · exact (congr_fun (congr_fun e₃ i) c :) · exact IH _ _ (h'' _) theorem M.bisim₀ {α : TypeVec n} (R : P.M α → P.M α → Prop) (h₀ : Equivalence R) @@ -284,7 +284,7 @@ theorem M.bisim' {α : TypeVec n} (R : P.M α → P.M α → Prop) induction Hr · rw [← Quot.factor_mk_eq R (Relation.EqvGen R) this] rwa [appendFun_comp_id, ← MvFunctor.map_map, ← MvFunctor.map_map, h] - all_goals aesop + all_goals simp_all theorem M.dest_map {α β : TypeVec n} (g : α ⟹ β) (x : P.M α) : M.dest P (g <$$> x) = (appendFun g fun x => g <$$> x) <$$> M.dest P x := by diff --git a/Mathlib/Data/PNat/Basic.lean b/Mathlib/Data/PNat/Basic.lean index 087e1dc61c787..5b28b50869c1a 100644 --- a/Mathlib/Data/PNat/Basic.lean +++ b/Mathlib/Data/PNat/Basic.lean @@ -56,12 +56,12 @@ theorem natPred_inj {m n : ℕ+} : m.natPred = n.natPred ↔ m = n := @[simp, norm_cast] lemma val_ofNat (n : ℕ) [NeZero n] : - ((no_index (OfNat.ofNat n) : ℕ+) : ℕ) = OfNat.ofNat n := + ((ofNat(n) : ℕ+) : ℕ) = OfNat.ofNat n := rfl @[simp] lemma mk_ofNat (n : ℕ) (h : 0 < n) : - @Eq ℕ+ (⟨no_index (OfNat.ofNat n), h⟩ : ℕ+) (haveI : NeZero n := ⟨h.ne'⟩; OfNat.ofNat n) := + @Eq ℕ+ (⟨ofNat(n), h⟩ : ℕ+) (haveI : NeZero n := ⟨h.ne'⟩; OfNat.ofNat n) := rfl end PNat @@ -182,17 +182,17 @@ theorem recOn_succ (n : ℕ+) {p : ℕ+ → Sort*} (p1 hp) : @[simp] theorem ofNat_le_ofNat {m n : ℕ} [NeZero m] [NeZero n] : - (no_index (OfNat.ofNat m) : ℕ+) ≤ no_index (OfNat.ofNat n) ↔ OfNat.ofNat m ≤ OfNat.ofNat n := + (ofNat(m) : ℕ+) ≤ ofNat(n) ↔ OfNat.ofNat m ≤ OfNat.ofNat n := .rfl @[simp] theorem ofNat_lt_ofNat {m n : ℕ} [NeZero m] [NeZero n] : - (no_index (OfNat.ofNat m) : ℕ+) < no_index (OfNat.ofNat n) ↔ OfNat.ofNat m < OfNat.ofNat n := + (ofNat(m) : ℕ+) < ofNat(n) ↔ OfNat.ofNat m < OfNat.ofNat n := .rfl @[simp] theorem ofNat_inj {m n : ℕ} [NeZero m] [NeZero n] : - (no_index (OfNat.ofNat m) : ℕ+) = no_index (OfNat.ofNat n) ↔ OfNat.ofNat m = OfNat.ofNat n := + (ofNat(m) : ℕ+) = ofNat(n) ↔ OfNat.ofNat m = OfNat.ofNat n := Subtype.mk_eq_mk @[simp, norm_cast] diff --git a/Mathlib/Data/Part.lean b/Mathlib/Data/Part.lean index 842a941412357..98ebf824ed437 100644 --- a/Mathlib/Data/Part.lean +++ b/Mathlib/Data/Part.lean @@ -397,7 +397,7 @@ theorem assert_pos {p : Prop} {f : p → Part α} (h : p) : assert p f = f h := simp only [h', mk.injEq, h, exists_prop_of_true, true_and] apply Function.hfunext · simp only [h, h', exists_prop_of_true] - · aesop + · simp theorem assert_neg {p : Prop} {f : p → Part α} (h : ¬p) : assert p f = none := by dsimp [assert, none]; congr diff --git a/Mathlib/Data/Prod/Basic.lean b/Mathlib/Data/Prod/Basic.lean index 37cb24c949727..d68eab7f60f81 100644 --- a/Mathlib/Data/Prod/Basic.lean +++ b/Mathlib/Data/Prod/Basic.lean @@ -251,10 +251,10 @@ theorem map_injective [Nonempty α] [Nonempty β] {f : α → γ} {g : β → δ ⟨fun a₁ a₂ ha => by inhabit β injection - @h (a₁, default) (a₂, default) (congr_arg (fun c : γ => Prod.mk c (g default)) ha : _), + @h (a₁, default) (a₂, default) (congr_arg (fun c : γ => Prod.mk c (g default)) ha :), fun b₁ b₂ hb => by inhabit α - injection @h (default, b₁) (default, b₂) (congr_arg (Prod.mk (f default)) hb : _)⟩, + injection @h (default, b₁) (default, b₂) (congr_arg (Prod.mk (f default)) hb :)⟩, fun h => h.1.prodMap h.2⟩ @[simp] diff --git a/Mathlib/Data/Rat/Cast/Defs.lean b/Mathlib/Data/Rat/Cast/Defs.lean index a5fc916543c00..5636ef6193774 100644 --- a/Mathlib/Data/Rat/Cast/Defs.lean +++ b/Mathlib/Data/Rat/Cast/Defs.lean @@ -114,8 +114,6 @@ theorem cast_intCast (n : ℤ) : ((n : ℚ) : α) = n := theorem cast_natCast (n : ℕ) : ((n : ℚ) : α) = n := by rw [← Int.cast_natCast, cast_intCast, Int.cast_natCast] -@[deprecated (since := "2024-03-21")] alias cast_coe_int := cast_intCast -@[deprecated (since := "2024-03-21")] alias cast_coe_nat := cast_natCast @[simp, norm_cast] lemma cast_ofNat (n : ℕ) [n.AtLeastTwo] : ((ofNat(n) : ℚ) : α) = (ofNat(n) : α) := by diff --git a/Mathlib/Data/Rat/Cast/Order.lean b/Mathlib/Data/Rat/Cast/Order.lean index 618e76fd32375..e5287b4093aaf 100644 --- a/Mathlib/Data/Rat/Cast/Order.lean +++ b/Mathlib/Data/Rat/Cast/Order.lean @@ -174,16 +174,16 @@ def castOrderEmbedding : ℚ≥0 ↪o K := section ofNat variable {n : ℕ} [n.AtLeastTwo] -@[simp] lemma cast_le_ofNat : (p : K) ≤ no_index (OfNat.ofNat n) ↔ p ≤ OfNat.ofNat n := by +@[simp] lemma cast_le_ofNat : (p : K) ≤ ofNat(n) ↔ p ≤ OfNat.ofNat n := by simp [← cast_le (K := K)] -@[simp] lemma ofNat_le_cast : no_index (OfNat.ofNat n) ≤ (p : K) ↔ OfNat.ofNat n ≤ p := by +@[simp] lemma ofNat_le_cast : ofNat(n) ≤ (p : K) ↔ OfNat.ofNat n ≤ p := by simp [← cast_le (K := K)] -@[simp] lemma cast_lt_ofNat : (p : K) < no_index (OfNat.ofNat n) ↔ p < OfNat.ofNat n := by +@[simp] lemma cast_lt_ofNat : (p : K) < ofNat(n) ↔ p < OfNat.ofNat n := by simp [← cast_lt (K := K)] -@[simp] lemma ofNat_lt_cast : no_index (OfNat.ofNat n) < (p : K) ↔ OfNat.ofNat n < p := by +@[simp] lemma ofNat_lt_cast : ofNat(n) < (p : K) ↔ OfNat.ofNat n < p := by simp [← cast_lt (K := K)] end ofNat diff --git a/Mathlib/Data/Rat/Defs.lean b/Mathlib/Data/Rat/Defs.lean index cb77cd4efc0c3..f848e1c3e14a3 100644 --- a/Mathlib/Data/Rat/Defs.lean +++ b/Mathlib/Data/Rat/Defs.lean @@ -167,24 +167,14 @@ theorem lift_binop_eq (f : ℚ → ℚ → ℚ) (f₁ : ℤ → ℤ → ℤ → attribute [simp] divInt_add_divInt -@[deprecated divInt_add_divInt (since := "2024-03-18")] -theorem add_def'' {a b c d : ℤ} (b0 : b ≠ 0) (d0 : d ≠ 0) : - a /. b + c /. d = (a * d + c * b) /. (b * d) := divInt_add_divInt _ _ b0 d0 - attribute [simp] neg_divInt lemma neg_def (q : ℚ) : -q = -q.num /. q.den := by rw [← neg_divInt, num_divInt_den] @[simp] lemma divInt_neg (n d : ℤ) : n /. -d = -n /. d := divInt_neg' .. -@[deprecated (since := "2024-03-18")] alias divInt_neg_den := divInt_neg - attribute [simp] divInt_sub_divInt -@[deprecated divInt_sub_divInt (since := "2024-03-18")] -lemma sub_def'' {a b c d : ℤ} (b0 : b ≠ 0) (d0 : d ≠ 0) : - a /. b - c /. d = (a * d - c * b) /. (b * d) := divInt_sub_divInt _ _ b0 d0 - @[simp] lemma divInt_mul_divInt' (n₁ d₁ n₂ d₂ : ℤ) : (n₁ /. d₁) * (n₂ /. d₂) = (n₁ * n₂) /. (d₁ * d₂) := by obtain rfl | h₁ := eq_or_ne d₁ 0 @@ -270,17 +260,11 @@ protected theorem add_assoc : a + b + c = a + (b + c) := protected lemma neg_add_cancel : -a + a = 0 := by simp [add_def, normalize_eq_mkRat, Int.neg_mul, Int.add_comm, ← Int.sub_eq_add_neg] -@[deprecated zero_divInt (since := "2024-03-18")] -lemma divInt_zero_one : 0 /. 1 = 0 := zero_divInt _ - @[simp] lemma divInt_one (n : ℤ) : n /. 1 = n := by simp [divInt, mkRat, normalize] @[simp] lemma mkRat_one (n : ℤ) : mkRat n 1 = n := by simp [mkRat_eq_divInt] lemma divInt_one_one : 1 /. 1 = 1 := by rw [divInt_one, intCast_one] -@[deprecated divInt_one (since := "2024-03-18")] -lemma divInt_neg_one_one : -1 /. 1 = -1 := by rw [divInt_one, intCast_neg, intCast_one] - protected theorem mul_assoc : a * b * c = a * (b * c) := numDenCasesOn' a fun n₁ d₁ h₁ => numDenCasesOn' b fun n₂ d₂ h₂ => diff --git a/Mathlib/Data/Rat/Lemmas.lean b/Mathlib/Data/Rat/Lemmas.lean index 739a18a81c1b1..e0ac6e8bf3883 100644 --- a/Mathlib/Data/Rat/Lemmas.lean +++ b/Mathlib/Data/Rat/Lemmas.lean @@ -125,10 +125,9 @@ theorem isSquare_natCast_iff {n : ℕ} : IsSquare (n : ℚ) ↔ IsSquare n := by theorem isSquare_intCast_iff {z : ℤ} : IsSquare (z : ℚ) ↔ IsSquare z := by simp_rw [isSquare_iff, intCast_num, intCast_den, IsSquare.one, and_true] --- See note [no_index around OfNat.ofNat] @[simp] theorem isSquare_ofNat_iff {n : ℕ} : - IsSquare (no_index (OfNat.ofNat n) : ℚ) ↔ IsSquare (OfNat.ofNat n : ℕ) := + IsSquare (ofNat(n) : ℚ) ↔ IsSquare (OfNat.ofNat n : ℕ) := isSquare_natCast_iff section Casts @@ -265,7 +264,7 @@ theorem inv_natCast_num (a : ℕ) : (a : ℚ)⁻¹.num = Int.sign a := inv_intCast_num a @[simp] -theorem inv_ofNat_num (a : ℕ) [a.AtLeastTwo] : (no_index (OfNat.ofNat a : ℚ))⁻¹.num = 1 := +theorem inv_ofNat_num (a : ℕ) [a.AtLeastTwo] : (ofNat(a) : ℚ)⁻¹.num = 1 := inv_natCast_num_of_pos (Nat.pos_of_neZero a) @[simp] @@ -301,7 +300,7 @@ theorem inv_natCast_den (a : ℕ) : (a : ℚ)⁻¹.den = if a = 0 then 1 else a @[simp] theorem inv_ofNat_den (a : ℕ) [a.AtLeastTwo] : - (no_index (OfNat.ofNat a : ℚ))⁻¹.den = OfNat.ofNat a := + (ofNat(a) : ℚ)⁻¹.den = OfNat.ofNat a := inv_natCast_den_of_pos (Nat.pos_of_neZero a) protected theorem «forall» {p : ℚ → Prop} : (∀ r, p r) ↔ ∀ a b : ℤ, p (a / b) := diff --git a/Mathlib/Data/Rat/Sqrt.lean b/Mathlib/Data/Rat/Sqrt.lean index 7c9fe65973ca4..5f6d920bde3b8 100644 --- a/Mathlib/Data/Rat/Sqrt.lean +++ b/Mathlib/Data/Rat/Sqrt.lean @@ -45,9 +45,8 @@ theorem sqrt_intCast (z : ℤ) : Rat.sqrt (z : ℚ) = Int.sqrt z := by theorem sqrt_natCast (n : ℕ) : Rat.sqrt (n : ℚ) = Nat.sqrt n := by rw [← Int.cast_natCast, sqrt_intCast, Int.sqrt_natCast, Int.cast_natCast] --- See note [no_index around OfNat.ofNat] @[simp] -theorem sqrt_ofNat (n : ℕ) : Rat.sqrt (no_index (OfNat.ofNat n) : ℚ) = Nat.sqrt (OfNat.ofNat n) := +theorem sqrt_ofNat (n : ℕ) : Rat.sqrt (ofNat(n) : ℚ) = Nat.sqrt (OfNat.ofNat n) := sqrt_natCast _ end Rat diff --git a/Mathlib/Data/Real/Hyperreal.lean b/Mathlib/Data/Real/Hyperreal.lean index fd849459ea9f9..4e7a601f04cdc 100644 --- a/Mathlib/Data/Real/Hyperreal.lean +++ b/Mathlib/Data/Real/Hyperreal.lean @@ -72,10 +72,9 @@ theorem coe_neg (x : ℝ) : ↑(-x) = (-x : ℝ*) := theorem coe_add (x y : ℝ) : ↑(x + y) = (x + y : ℝ*) := rfl --- See note [no_index around OfNat.ofNat] @[simp, norm_cast] theorem coe_ofNat (n : ℕ) [n.AtLeastTwo] : - ((no_index (OfNat.ofNat n : ℝ)) : ℝ*) = OfNat.ofNat n := + ((ofNat(n) : ℝ) : ℝ*) = OfNat.ofNat n := rfl @[simp, norm_cast] diff --git a/Mathlib/Data/Real/Sqrt.lean b/Mathlib/Data/Real/Sqrt.lean index 5b1193b70b404..14d02f76715e2 100644 --- a/Mathlib/Data/Real/Sqrt.lean +++ b/Mathlib/Data/Real/Sqrt.lean @@ -62,11 +62,6 @@ lemma sqrt_le_iff_le_sq : sqrt x ≤ y ↔ x ≤ y ^ 2 := sqrt.to_galoisConnecti lemma le_sqrt_iff_sq_le : x ≤ sqrt y ↔ x ^ 2 ≤ y := (sqrt.symm.to_galoisConnection _ _).symm -@[deprecated (since := "2024-02-14")] alias sqrt_le_sqrt_iff := sqrt_le_sqrt -@[deprecated (since := "2024-02-14")] alias sqrt_lt_sqrt_iff := sqrt_lt_sqrt -@[deprecated (since := "2024-02-14")] alias sqrt_le_iff := sqrt_le_iff_le_sq -@[deprecated (since := "2024-02-14")] alias le_sqrt_iff := le_sqrt_iff_sq_le -@[deprecated (since := "2024-02-14")] alias sqrt_eq_iff_sq_eq := sqrt_eq_iff_eq_sq @[simp] lemma sqrt_eq_zero : sqrt x = 0 ↔ x = 0 := by simp [sqrt_eq_iff_eq_sq] diff --git a/Mathlib/Data/Set/Basic.lean b/Mathlib/Data/Set/Basic.lean index 64f61f2b25e36..0a7d5df2df1b0 100644 --- a/Mathlib/Data/Set/Basic.lean +++ b/Mathlib/Data/Set/Basic.lean @@ -514,7 +514,6 @@ theorem subset_eq_empty {s t : Set α} (h : t ⊆ s) (e : s = ∅) : t = ∅ := theorem forall_mem_empty {p : α → Prop} : (∀ x ∈ (∅ : Set α), p x) ↔ True := iff_true_intro fun _ => False.elim -@[deprecated (since := "2024-03-23")] alias ball_empty_iff := forall_mem_empty instance (α : Type u) : IsEmpty.{u + 1} (↥(∅ : Set α)) := ⟨fun x => x.2⟩ @@ -843,10 +842,6 @@ theorem union_inter_distrib_left (s t u : Set α) : s ∪ t ∩ u = (s ∪ t) theorem inter_union_distrib_right (s t u : Set α) : s ∩ t ∪ u = (s ∪ u) ∩ (t ∪ u) := sup_inf_right _ _ _ -@[deprecated (since := "2024-03-22")] alias inter_distrib_left := inter_union_distrib_left -@[deprecated (since := "2024-03-22")] alias inter_distrib_right := union_inter_distrib_right -@[deprecated (since := "2024-03-22")] alias union_distrib_left := union_inter_distrib_left -@[deprecated (since := "2024-03-22")] alias union_distrib_right := inter_union_distrib_right theorem union_union_distrib_left (s t u : Set α) : s ∪ (t ∪ u) = s ∪ t ∪ (s ∪ u) := sup_sup_distrib_left _ _ _ @@ -979,12 +974,10 @@ theorem forall_insert_of_forall {P : α → Prop} {a : α} {s : Set α} (H : ∀ theorem exists_mem_insert {P : α → Prop} {a : α} {s : Set α} : (∃ x ∈ insert a s, P x) ↔ (P a ∨ ∃ x ∈ s, P x) := by simp [mem_insert_iff, or_and_right, exists_and_left, exists_or] -@[deprecated (since := "2024-03-23")] alias bex_insert_iff := exists_mem_insert theorem forall_mem_insert {P : α → Prop} {a : α} {s : Set α} : (∀ x ∈ insert a s, P x) ↔ P a ∧ ∀ x ∈ s, P x := forall₂_or_left.trans <| and_congr_left' forall_eq -@[deprecated (since := "2024-03-23")] alias ball_insert_iff := forall_mem_insert /-- Inserting an element to a set is equivalent to the option type. -/ def subtypeInsertEquivOption diff --git a/Mathlib/Data/Set/Finite/Lattice.lean b/Mathlib/Data/Set/Finite/Lattice.lean index 85c9d8dbf4beb..0f87a472988fd 100644 --- a/Mathlib/Data/Set/Finite/Lattice.lean +++ b/Mathlib/Data/Set/Finite/Lattice.lean @@ -152,6 +152,35 @@ theorem Finite.iUnion {ι : Type*} {s : ι → Set α} {t : Set ι} (ht : t.Fini · rw [he i hi, mem_empty_iff_false] at hx contradiction +/-- An indexed union of pairwise disjoint sets is finite iff all sets are finite, and all but +finitely many are empty. -/ +lemma finite_iUnion_iff {ι : Type*} {s : ι → Set α} (hs : Pairwise fun i j ↦ Disjoint (s i) (s j)) : + (⋃ i, s i).Finite ↔ (∀ i, (s i).Finite) ∧ {i | (s i).Nonempty}.Finite where + mp h := by + refine ⟨fun i ↦ h.subset <| subset_iUnion _ _, ?_⟩ + let u (i : {i | (s i).Nonempty}) : ⋃ i, s i := ⟨i.2.choose, mem_iUnion.2 ⟨i.1, i.2.choose_spec⟩⟩ + have u_inj : Function.Injective u := by + rintro ⟨i, hi⟩ ⟨j, hj⟩ hij + ext + refine hs.eq <| not_disjoint_iff.2 ⟨u ⟨i, hi⟩, hi.choose_spec, ?_⟩ + rw [hij] + exact hj.choose_spec + have : Finite (⋃ i, s i) := h + exact .of_injective u u_inj + mpr h := h.2.iUnion (fun _ _ ↦ h.1 _) (by simp [not_nonempty_iff_eq_empty]) + +@[simp] lemma finite_iUnion_of_subsingleton {ι : Sort*} [Subsingleton ι] {s : ι → Set α} : + (⋃ i, s i).Finite ↔ ∀ i, (s i).Finite := by + rw [← iUnion_plift_down, finite_iUnion_iff _root_.Subsingleton.pairwise] + simp [PLift.forall, Finite.of_subsingleton] + +/-- An indexed union of pairwise disjoint sets is finite iff all sets are finite, and all but +finitely many are empty. -/ +lemma PairwiseDisjoint.finite_biUnion_iff {f : β → Set α} {s : Set β} (hs : s.PairwiseDisjoint f) : + (⋃ i ∈ s, f i).Finite ↔ (∀ i ∈ s, (f i).Finite) ∧ {i ∈ s | (f i).Nonempty}.Finite := by + rw [finite_iUnion_iff (by aesop (add unfold safe [Pairwise, PairwiseDisjoint, Set.Pairwise]))] + simp + section preimage variable {f : α → β} {s : Set β} diff --git a/Mathlib/Data/Set/Function.lean b/Mathlib/Data/Set/Function.lean index f2c6f67261d56..c29e123ed9246 100644 --- a/Mathlib/Data/Set/Function.lean +++ b/Mathlib/Data/Set/Function.lean @@ -529,6 +529,10 @@ theorem _root_.Function.Injective.injOn_range (h : Injective (g ∘ f)) : InjOn rintro _ ⟨x, rfl⟩ _ ⟨y, rfl⟩ H exact congr_arg f (h H) +theorem _root_.Set.InjOn.injective_iff (s : Set β) (h : InjOn g s) (hs : range f ⊆ s) : + Injective (g ∘ f) ↔ Injective f := + ⟨(·.of_comp), fun h _ ↦ by aesop⟩ + theorem injOn_iff_injective : InjOn f s ↔ Injective (s.restrict f) := ⟨fun H a b h => Subtype.eq <| H a.2 b.2 h, fun H a as b bs h => congr_arg Subtype.val <| @H ⟨a, as⟩ ⟨b, bs⟩ h⟩ diff --git a/Mathlib/Data/Set/Image.lean b/Mathlib/Data/Set/Image.lean index a98696d44ae14..a36c4d6171444 100644 --- a/Mathlib/Data/Set/Image.lean +++ b/Mathlib/Data/Set/Image.lean @@ -186,11 +186,6 @@ variable {f : α → β} {s t : Set α} -- Porting note: `Set.image` is already defined in `Data.Set.Defs` -@[deprecated mem_image (since := "2024-03-23")] -theorem mem_image_iff_bex {f : α → β} {s : Set α} {y : β} : - y ∈ f '' s ↔ ∃ (x : _) (_ : x ∈ s), f x = y := - bex_def.symm - theorem image_eta (f : α → β) : f '' s = (fun x => f x) '' s := rfl @@ -208,18 +203,6 @@ theorem forall_mem_image {f : α → β} {s : Set α} {p : β → Prop} : theorem exists_mem_image {f : α → β} {s : Set α} {p : β → Prop} : (∃ y ∈ f '' s, p y) ↔ ∃ x ∈ s, p (f x) := by simp -@[deprecated (since := "2024-02-21")] alias ball_image_iff := forall_mem_image -@[deprecated (since := "2024-02-21")] alias bex_image_iff := exists_mem_image -@[deprecated (since := "2024-02-21")] alias ⟨_, ball_image_of_ball⟩ := forall_mem_image - -@[deprecated forall_mem_image (since := "2024-02-21")] -theorem mem_image_elim {f : α → β} {s : Set α} {C : β → Prop} (h : ∀ x : α, x ∈ s → C (f x)) : - ∀ {y : β}, y ∈ f '' s → C y := forall_mem_image.2 h _ - -@[deprecated forall_mem_image (since := "2024-02-21")] -theorem mem_image_elim_on {f : α → β} {s : Set α} {C : β → Prop} {y : β} (h_y : y ∈ f '' s) - (h : ∀ x : α, x ∈ s → C (f x)) : C y := forall_mem_image.2 h _ h_y - -- Porting note: used to be `safe` @[congr] theorem image_congr {f g : α → β} {s : Set α} (h : ∀ a ∈ s, f a = g a) : f '' s = g '' s := by @@ -574,7 +557,6 @@ variable {f : ι → α} {s t : Set α} theorem forall_mem_range {p : α → Prop} : (∀ a ∈ range f, p a) ↔ ∀ i, p (f i) := by simp -@[deprecated (since := "2024-02-21")] alias forall_range_iff := forall_mem_range theorem forall_subtype_range_iff {p : range f → Prop} : (∀ a : range f, p a) ↔ ∀ i, p ⟨f i, mem_range_self _⟩ := @@ -584,9 +566,6 @@ theorem forall_subtype_range_iff {p : range f → Prop} : theorem exists_range_iff {p : α → Prop} : (∃ a ∈ range f, p a) ↔ ∃ i, p (f i) := by simp -@[deprecated (since := "2024-03-10")] -alias exists_range_iff' := exists_range_iff - theorem exists_subtype_range_iff {p : range f → Prop} : (∃ a : range f, p a) ↔ ∃ i, p ⟨f i, mem_range_self _⟩ := ⟨fun ⟨⟨a, i, hi⟩, ha⟩ => by diff --git a/Mathlib/Data/Setoid/Basic.lean b/Mathlib/Data/Setoid/Basic.lean index 3599ac45fefc2..48dc6951f8870 100644 --- a/Mathlib/Data/Setoid/Basic.lean +++ b/Mathlib/Data/Setoid/Basic.lean @@ -89,6 +89,8 @@ theorem trans' (r : Setoid α) : ∀ {x y z}, r x y → r y z → r x z := r.ise theorem comm' (s : Setoid α) {x y} : s x y ↔ s y x := ⟨s.symm', s.symm'⟩ +open scoped Function -- required for scoped `on` notation + /-- The kernel of a function is an equivalence relation. -/ def ker (f : α → β) : Setoid α := ⟨(· = ·) on f, eq_equivalence.comap f⟩ diff --git a/Mathlib/Data/Setoid/Partition.lean b/Mathlib/Data/Setoid/Partition.lean index 7781e1a597991..c0909554e4007 100644 --- a/Mathlib/Data/Setoid/Partition.lean +++ b/Mathlib/Data/Setoid/Partition.lean @@ -310,6 +310,8 @@ structure IndexedPartition {ι α : Type*} (s : ι → Set α) where /-- membership invariance for `index`-/ mem_index : ∀ x, x ∈ s (index x) +open scoped Function -- required for scoped `on` notation + /-- The non-constructive constructor for `IndexedPartition`. -/ noncomputable def IndexedPartition.mk' {ι α : Type*} (s : ι → Set α) (dis : Pairwise (Disjoint on s)) (nonempty : ∀ i, (s i).Nonempty) diff --git a/Mathlib/Data/Sum/Basic.lean b/Mathlib/Data/Sum/Basic.lean index 9631344ff2cbf..5dd68ca13d171 100644 --- a/Mathlib/Data/Sum/Basic.lean +++ b/Mathlib/Data/Sum/Basic.lean @@ -198,8 +198,8 @@ open Function theorem map_injective {f : α → γ} {g : β → δ} : Injective (Sum.map f g) ↔ Injective f ∧ Injective g := ⟨fun h => - ⟨fun a₁ a₂ ha => inl_injective <| @h (inl a₁) (inl a₂) (congr_arg inl ha : _), fun b₁ b₂ hb => - inr_injective <| @h (inr b₁) (inr b₂) (congr_arg inr hb : _)⟩, + ⟨fun a₁ a₂ ha => inl_injective <| @h (inl a₁) (inl a₂) (congr_arg inl ha :), fun b₁ b₂ hb => + inr_injective <| @h (inr b₁) (inr b₂) (congr_arg inr hb :)⟩, fun h => h.1.sum_map h.2⟩ @[simp] diff --git a/Mathlib/Data/Sum/Interval.lean b/Mathlib/Data/Sum/Interval.lean index f376aa6460b2f..9d51847e6bfbb 100644 --- a/Mathlib/Data/Sum/Interval.lean +++ b/Mathlib/Data/Sum/Interval.lean @@ -217,10 +217,10 @@ instance instLocallyFiniteOrder : LocallyFiniteOrder (α ⊕ β) where finsetIco := sumLift₂ Ico Ico finsetIoc := sumLift₂ Ioc Ioc finsetIoo := sumLift₂ Ioo Ioo - finset_mem_Icc := by rintro (a | a) (b | b) (x | x) <;> simp - finset_mem_Ico := by rintro (a | a) (b | b) (x | x) <;> simp - finset_mem_Ioc := by rintro (a | a) (b | b) (x | x) <;> simp - finset_mem_Ioo := by rintro (a | a) (b | b) (x | x) <;> simp + finset_mem_Icc := by simp + finset_mem_Ico := by simp + finset_mem_Ioc := by simp + finset_mem_Ioo := by simp variable (a₁ a₂ : α) (b₁ b₂ : β) diff --git a/Mathlib/Data/Sym/Basic.lean b/Mathlib/Data/Sym/Basic.lean index 2733966c42fc7..24af84e1ec2d4 100644 --- a/Mathlib/Data/Sym/Basic.lean +++ b/Mathlib/Data/Sym/Basic.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Kyle Miller. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kyle Miller -/ -import Mathlib.Data.Multiset.Basic +import Mathlib.Algebra.Order.Group.Multiset import Mathlib.Data.Vector.Basic import Mathlib.Data.Setoid.Basic import Mathlib.Tactic.ApplyFun diff --git a/Mathlib/Data/TypeVec.lean b/Mathlib/Data/TypeVec.lean index 88be36a0d3c00..7c4cbef504339 100644 --- a/Mathlib/Data/TypeVec.lean +++ b/Mathlib/Data/TypeVec.lean @@ -376,7 +376,7 @@ theorem const_nil {β} (x : β) (α : TypeVec 0) : TypeVec.const x α = nilFun : @[typevec] theorem repeat_eq_append1 {β} {n} (α : TypeVec n) : - repeatEq (α ::: β) = splitFun (α := (α ⊗ α) ::: _ ) + repeatEq (α ::: β) = splitFun (α := (α ⊗ α) ::: _) (α' := («repeat» n Prop) ::: _) (repeatEq α) (uncurry Eq) := by induction n <;> rfl diff --git a/Mathlib/Data/W/Basic.lean b/Mathlib/Data/W/Basic.lean index 7c1f339cc49d0..83c3b89cad539 100644 --- a/Mathlib/Data/W/Basic.lean +++ b/Mathlib/Data/W/Basic.lean @@ -82,7 +82,7 @@ theorem elim_injective (γ : Type*) (fγ : (Σa : α, β a → γ) → γ) | ⟨a₁, f₁⟩, ⟨a₂, f₂⟩, h => by obtain ⟨rfl, h⟩ := Sigma.mk.inj_iff.mp (fγ_injective h) congr with x - exact elim_injective γ fγ fγ_injective (congr_fun (eq_of_heq h) x : _) + exact elim_injective γ fγ fγ_injective (congr_fun (eq_of_heq h) x :) instance [hα : IsEmpty α] : IsEmpty (WType β) := ⟨fun w => WType.recOn w (IsEmpty.elim hα)⟩ diff --git a/Mathlib/Data/ZMod/Basic.lean b/Mathlib/Data/ZMod/Basic.lean index 33ba9dfef1f77..e6ccd28f30073 100644 --- a/Mathlib/Data/ZMod/Basic.lean +++ b/Mathlib/Data/ZMod/Basic.lean @@ -1046,7 +1046,7 @@ theorem neg_eq_self_iff {n : ℕ} (a : ZMod n) : -a = a ↔ a = 0 ∨ 2 * a.val refine (a.val_lt.not_le <| Nat.le_of_mul_le_mul_left ?_ zero_lt_two).elim rw [he, mul_comm] apply Nat.mul_le_mul_left - erw [Nat.succ_le_succ_iff, Nat.succ_le_succ_iff]; simp + simp · rintro (rfl | h) · rw [val_zero, mul_zero] apply dvd_zero diff --git a/Mathlib/Data/ZMod/IntUnitsPower.lean b/Mathlib/Data/ZMod/IntUnitsPower.lean index 61cc3341fdfc7..80e8edf23f8c4 100644 --- a/Mathlib/Data/ZMod/IntUnitsPower.lean +++ b/Mathlib/Data/ZMod/IntUnitsPower.lean @@ -71,9 +71,8 @@ example : Int.instUnitsPow = DivInvMonoid.toZPow := rfl change ((n : R) • Additive.ofMul u).toMul = _ rw [Nat.cast_smul_eq_nsmul, toMul_nsmul, toMul_ofMul] --- See note [no_index around OfNat.ofNat] lemma uzpow_coe_nat (s : ℤˣ) (n : ℕ) [n.AtLeastTwo] : - s ^ (no_index (OfNat.ofNat n : R)) = s ^ (no_index (OfNat.ofNat n : ℕ)) := + s ^ (ofNat(n) : R) = s ^ (ofNat(n) : ℕ) := uzpow_natCast _ _ @[simp] lemma one_uzpow (x : R) : (1 : ℤˣ) ^ x = 1 := diff --git a/Mathlib/Data/ZMod/Quotient.lean b/Mathlib/Data/ZMod/Quotient.lean index bcbc87a1d2aa0..8f05b3b2992d5 100644 --- a/Mathlib/Data/ZMod/Quotient.lean +++ b/Mathlib/Data/ZMod/Quotient.lean @@ -63,6 +63,7 @@ end Int noncomputable section ChineseRemainder open Ideal +open scoped Function in -- required for scoped `on` notation /-- The **Chinese remainder theorem**, elementary version for `ZMod`. See also `Mathlib.Data.ZMod.Basic` for versions involving only two numbers. -/ def ZMod.prodEquivPi {ι : Type*} [Fintype ι] (a : ι → ℕ) @@ -142,9 +143,6 @@ theorem orbitZPowersEquiv_symm_apply (k : ZMod (minimalPeriod (a • ·) b)) : (⟨a, mem_zpowers a⟩ : zpowers a) ^ (cast k : ℤ) • ⟨b, mem_orbit_self b⟩ := rfl -@[deprecated (since := "2024-02-21")] -alias _root_.AddAction.orbit_zmultiples_equiv_symm_apply := orbitZMultiplesEquiv_symm_apply - theorem orbitZPowersEquiv_symm_apply' (k : ℤ) : (orbitZPowersEquiv a b).symm k = (⟨a, mem_zpowers a⟩ : zpowers a) ^ k • ⟨b, mem_orbit_self b⟩ := by diff --git a/Mathlib/Deprecated/Cardinal/PartENat.lean b/Mathlib/Deprecated/Cardinal/PartENat.lean index e2c263a8cfb2e..47e21caf7cb31 100644 --- a/Mathlib/Deprecated/Cardinal/PartENat.lean +++ b/Mathlib/Deprecated/Cardinal/PartENat.lean @@ -48,9 +48,6 @@ theorem toPartENat_eq_top {c : Cardinal} : theorem toPartENat_apply_of_aleph0_le {c : Cardinal} (h : ℵ₀ ≤ c) : toPartENat c = ⊤ := congr_arg PartENat.ofENat (toENat_eq_top.2 h) -@[deprecated (since := "2024-02-15")] -alias toPartENat_cast := toPartENat_natCast - @[simp] theorem mk_toPartENat_of_infinite [h : Infinite α] : toPartENat #α = ⊤ := toPartENat_apply_of_aleph0_le (infinite_iff.1 h) @@ -62,7 +59,6 @@ theorem aleph0_toPartENat : toPartENat ℵ₀ = ⊤ := theorem toPartENat_surjective : Surjective toPartENat := fun x => PartENat.casesOn x ⟨ℵ₀, toPartENat_apply_of_aleph0_le le_rfl⟩ fun n => ⟨n, toPartENat_natCast n⟩ -@[deprecated (since := "2024-02-15")] alias toPartENat_eq_top_iff_le_aleph0 := toPartENat_eq_top theorem toPartENat_strictMonoOn : StrictMonoOn toPartENat (Set.Iic ℵ₀) := PartENat.withTopOrderIso.symm.strictMono.comp_strictMonoOn toENat_strictMonoOn diff --git a/Mathlib/Deprecated/Submonoid.lean b/Mathlib/Deprecated/Submonoid.lean index 2e9618d300fb9..104e7376fd117 100644 --- a/Mathlib/Deprecated/Submonoid.lean +++ b/Mathlib/Deprecated/Submonoid.lean @@ -184,7 +184,6 @@ theorem IsSubmonoid.pow_mem {a : M} (hs : IsSubmonoid s) (h : a ∈ s) : ∀ {n the `AddSubmonoid`."] theorem IsSubmonoid.powers_subset {a : M} (hs : IsSubmonoid s) (h : a ∈ s) : powers a ⊆ s := fun _ ⟨_, hx⟩ => hx ▸ hs.pow_mem h -@[deprecated (since := "2024-02-21")] alias IsSubmonoid.power_subset := IsSubmonoid.powers_subset end powers diff --git a/Mathlib/Dynamics/FixedPoints/Basic.lean b/Mathlib/Dynamics/FixedPoints/Basic.lean index 78081cb348764..ed9a9b65367ad 100644 --- a/Mathlib/Dynamics/FixedPoints/Basic.lean +++ b/Mathlib/Dynamics/FixedPoints/Basic.lean @@ -35,7 +35,7 @@ open Function (Commute) /-- Every point is a fixed point of `id`. -/ theorem isFixedPt_id (x : α) : IsFixedPt id x := - (rfl : _) + (rfl :) namespace IsFixedPt diff --git a/Mathlib/FieldTheory/Adjoin.lean b/Mathlib/FieldTheory/Adjoin.lean index 80352130ce0b4..e0d6ba6cd6a31 100644 --- a/Mathlib/FieldTheory/Adjoin.lean +++ b/Mathlib/FieldTheory/Adjoin.lean @@ -928,7 +928,7 @@ theorem exists_finset_of_mem_supr' {ι : Type*} {f : ι → IntermediateField F theorem exists_finset_of_mem_supr'' {ι : Type*} {f : ι → IntermediateField F E} (h : ∀ i, Algebra.IsAlgebraic F (f i)) {x : E} (hx : x ∈ ⨆ i, f i) : - ∃ s : Finset (Σ i, f i), x ∈ ⨆ i ∈ s, adjoin F ((minpoly F (i.2 : _)).rootSet E) := by + ∃ s : Finset (Σ i, f i), x ∈ ⨆ i ∈ s, adjoin F ((minpoly F (i.2 :)).rootSet E) := by -- Porting note: writing `fun i x1 hx1 => ...` does not work. refine exists_finset_of_mem_iSup (SetLike.le_def.mp (iSup_le (fun i => ?_)) hx) intro x1 hx1 @@ -1099,7 +1099,7 @@ noncomputable def adjoinRootEquivAdjoin (h : IsIntegral F α) : AlgEquiv.ofBijective (AdjoinRoot.liftHom (minpoly F α) (AdjoinSimple.gen F α) (aeval_gen_minpoly F α)) (by - set f := AdjoinRoot.lift _ _ (aeval_gen_minpoly F α : _) + set f := AdjoinRoot.lift _ _ (aeval_gen_minpoly F α :) haveI := Fact.mk (minpoly.irreducible h) constructor · exact RingHom.injective f @@ -1163,7 +1163,7 @@ theorem isAlgebraic_adjoin_simple {x : L} (hx : IsIntegral K x) : Algebra.IsAlge @[stacks 09GN] theorem adjoin.finrank {x : L} (hx : IsIntegral K x) : Module.finrank K K⟮x⟯ = (minpoly K x).natDegree := by - rw [PowerBasis.finrank (adjoin.powerBasis hx : _)] + rw [PowerBasis.finrank (adjoin.powerBasis hx :)] rfl /-- If `K / E / F` is a field extension tower, `S ⊂ K` is such that `F(S) = K`, diff --git a/Mathlib/FieldTheory/Galois/Basic.lean b/Mathlib/FieldTheory/Galois/Basic.lean index 5b30988548f8e..1213496da816d 100644 --- a/Mathlib/FieldTheory/Galois/Basic.lean +++ b/Mathlib/FieldTheory/Galois/Basic.lean @@ -286,11 +286,8 @@ def galoisCoinsertionIntermediateFieldSubgroup [FiniteDimensional F E] [IsGalois GaloisCoinsertion (OrderDual.toDual ∘ (IntermediateField.fixingSubgroup : IntermediateField F E → Subgroup (E ≃ₐ[F] E))) ((IntermediateField.fixedField : Subgroup (E ≃ₐ[F] E) → IntermediateField F E) ∘ - OrderDual.toDual) where - choice H _ := IntermediateField.fixedField H - gc K H := (IntermediateField.le_iff_le H K).symm - u_l_le K := le_of_eq (fixedField_fixingSubgroup K) - choice_eq _ _ := rfl + OrderDual.toDual) := + OrderIso.toGaloisCoinsertion intermediateFieldEquivSubgroup end IsGalois diff --git a/Mathlib/FieldTheory/Galois/Infinite.lean b/Mathlib/FieldTheory/Galois/Infinite.lean index c97aeaad84aa5..d93b6f9fec8dc 100644 --- a/Mathlib/FieldTheory/Galois/Infinite.lean +++ b/Mathlib/FieldTheory/Galois/Infinite.lean @@ -180,19 +180,15 @@ lemma fixingSubgroup_fixedField (H : ClosedSubgroup (K ≃ₐ[k] K)) [IsGalois k /-- The Galois correspondence from intermediate fields to closed subgroups. -/ def IntermediateFieldEquivClosedSubgroup [IsGalois k K] : IntermediateField k K ≃o (ClosedSubgroup (K ≃ₐ[k] K))ᵒᵈ where - toFun := fun L => - { L.fixingSubgroup with - isClosed' := fixingSubgroup_isClosed L } - invFun := fun H => IntermediateField.fixedField H.1 - left_inv := fun L => fixedField_fixingSubgroup L - right_inv := by - intro H + toFun L := ⟨L.fixingSubgroup, fixingSubgroup_isClosed L⟩ + invFun H := IntermediateField.fixedField H.1 + left_inv L := fixedField_fixingSubgroup L + right_inv H := by simp_rw [fixingSubgroup_fixedField H] rfl - map_rel_iff' := by - intro L₁ L₂ - show L₁.fixingSubgroup ≥ L₂.fixingSubgroup ↔ L₁ ≤ L₂ - rw [← fixedField_fixingSubgroup L₂, IntermediateField.le_iff_le, fixedField_fixingSubgroup L₂] + map_rel_iff' {K L} := by + rw [← fixedField_fixingSubgroup L, IntermediateField.le_iff_le, fixedField_fixingSubgroup L] + rfl /-- The Galois correspondence as a `GaloisInsertion` -/ def GaloisInsertionIntermediateFieldClosedSubgroup [IsGalois k K] : @@ -205,16 +201,14 @@ def GaloisInsertionIntermediateFieldClosedSubgroup [IsGalois k K] : /-- The Galois correspondence as a `GaloisCoinsertion` -/ def GaloisCoinsertionIntermediateFieldSubgroup [IsGalois k K] : GaloisCoinsertion (OrderDual.toDual ∘ fun (E : IntermediateField k K) ↦ E.fixingSubgroup) - ((fun (H : Subgroup (K ≃ₐ[k] K)) ↦ IntermediateField.fixedField H) ∘ - OrderDual.toDual) where + ((fun (H : Subgroup (K ≃ₐ[k] K)) ↦ IntermediateField.fixedField H) ∘ OrderDual.toDual) where choice H _ := IntermediateField.fixedField H gc E H := (IntermediateField.le_iff_le H E).symm u_l_le K := le_of_eq (fixedField_fixingSubgroup K) choice_eq _ _ := rfl theorem isOpen_iff_finite (L : IntermediateField k K) [IsGalois k K] : - IsOpen (IntermediateFieldEquivClosedSubgroup L).carrier ↔ - (FiniteDimensional k L) := by + IsOpen L.fixingSubgroup.carrier ↔ FiniteDimensional k L := by refine ⟨fun h ↦ ?_, fun h ↦ IntermediateField.fixingSubgroup_isOpen L⟩ have : (IntermediateFieldEquivClosedSubgroup.toFun L).carrier ∈ nhds 1 := IsOpen.mem_nhds h (congrFun rfl) @@ -234,14 +228,12 @@ theorem isOpen_iff_finite (L : IntermediateField k K) [IsGalois k K] : exact FiniteDimensional.left k L L'.1 theorem normal_iff_isGalois (L : IntermediateField k K) [IsGalois k K] : - Subgroup.Normal (IntermediateFieldEquivClosedSubgroup L).1 ↔ - IsGalois k L := by + L.fixingSubgroup.Normal ↔ IsGalois k L := by refine ⟨fun h ↦ ?_, fun h ↦ ?_⟩ · let f : L → IntermediateField k K := fun x => IntermediateField.lift <| - IntermediateField.fixedField <| Subgroup.map (restrictNormalHom - (adjoin k {x.1})) L.fixingSubgroup - have h' (x : K) : (Subgroup.map (restrictNormalHom - (adjoin k {x})) L.fixingSubgroup).Normal := + IntermediateField.fixedField <| Subgroup.map (restrictNormalHom (adjoin k {x.1})) + L.fixingSubgroup + have h' (x : K) : (Subgroup.map (restrictNormalHom (adjoin k {x})) L.fixingSubgroup).Normal := Subgroup.Normal.map h (restrictNormalHom (adjoin k {x})) (restrictNormalHom_surjective K) have n' (l : L) : IsGalois k (IntermediateField.fixedField <| Subgroup.map (restrictNormalHom (adjoin k {l.1})) L.fixingSubgroup) := by @@ -267,4 +259,9 @@ theorem normal_iff_isGalois (L : IntermediateField k K) [IsGalois k K] : · simpa only [IntermediateFieldEquivClosedSubgroup, RelIso.coe_fn_mk, Equiv.coe_fn_mk, ← L.restrictNormalHom_ker] using MonoidHom.normal_ker (restrictNormalHom L) +theorem isOpen_and_normal_iff_finite_and_isGalois (L : IntermediateField k K) [IsGalois k K] : + IsOpen L.fixingSubgroup.carrier ∧ L.fixingSubgroup.Normal ↔ + FiniteDimensional k L ∧ IsGalois k L := by + rw [isOpen_iff_finite, normal_iff_isGalois] + end InfiniteGalois diff --git a/Mathlib/FieldTheory/IsAlgClosed/Basic.lean b/Mathlib/FieldTheory/IsAlgClosed/Basic.lean index a8d5a85385e1b..156df9e238ca5 100644 --- a/Mathlib/FieldTheory/IsAlgClosed/Basic.lean +++ b/Mathlib/FieldTheory/IsAlgClosed/Basic.lean @@ -382,7 +382,7 @@ noncomputable def equivOfAlgebraic' [Nontrivial S] [NoZeroSMulDivisors R S] letI : NoZeroSMulDivisors R L := NoZeroSMulDivisors.of_algebraMap_injective <| by rw [IsScalarTower.algebraMap_eq R S L] exact (Function.Injective.comp (NoZeroSMulDivisors.algebraMap_injective S L) - (NoZeroSMulDivisors.algebraMap_injective R S) : _) + (NoZeroSMulDivisors.algebraMap_injective R S) :) letI : IsAlgClosure R L := { isAlgClosed := IsAlgClosure.isAlgClosed S isAlgebraic := ‹_› } diff --git a/Mathlib/FieldTheory/Normal.lean b/Mathlib/FieldTheory/Normal.lean index 12486d2a18034..c209741701297 100644 --- a/Mathlib/FieldTheory/Normal.lean +++ b/Mathlib/FieldTheory/Normal.lean @@ -142,7 +142,7 @@ instance normal_iSup {ι : Type*} (t : ι → IntermediateField F K) [h : ∀ i, Normal F (⨆ i, t i : IntermediateField F K) := by refine { toIsAlgebraic := isAlgebraic_iSup fun i => (h i).1, splits' := fun x => ?_ } obtain ⟨s, hx⟩ := exists_finset_of_mem_supr'' (fun i => (h i).1) x.2 - let E : IntermediateField F K := ⨆ i ∈ s, adjoin F ((minpoly F (i.2 : _)).rootSet K) + let E : IntermediateField F K := ⨆ i ∈ s, adjoin F ((minpoly F (i.2 :)).rootSet K) have hF : Normal F E := by haveI : IsSplittingField F E (∏ i ∈ s, minpoly F i.snd) := by refine isSplittingField_iSup ?_ fun i _ => adjoin_rootSet_isSplittingField ?_ diff --git a/Mathlib/FieldTheory/PrimitiveElement.lean b/Mathlib/FieldTheory/PrimitiveElement.lean index 6b12f4601768f..b8174c3c41d64 100644 --- a/Mathlib/FieldTheory/PrimitiveElement.lean +++ b/Mathlib/FieldTheory/PrimitiveElement.lean @@ -286,9 +286,6 @@ theorem FiniteDimensional.of_finite_intermediateField rw [htop] at hfin exact topEquiv.toLinearEquiv.finiteDimensional -@[deprecated (since := "2024-02-02")] -alias finiteDimensional_of_finite_intermediateField := FiniteDimensional.of_finite_intermediateField - theorem exists_primitive_element_of_finite_intermediateField [Finite (IntermediateField F E)] (K : IntermediateField F E) : ∃ α : E, F⟮α⟯ = K := by haveI := FiniteDimensional.of_finite_intermediateField F E @@ -307,9 +304,6 @@ theorem FiniteDimensional.of_exists_primitive_element [Algebra.IsAlgebraic F E] rw [hprim] at hfin exact topEquiv.toLinearEquiv.finiteDimensional -@[deprecated (since := "2024-02-02")] -alias finiteDimensional_of_exists_primitive_element := FiniteDimensional.of_exists_primitive_element - -- A finite simple extension has only finitely many intermediate fields theorem finite_intermediateField_of_exists_primitive_element [Algebra.IsAlgebraic F E] (h : ∃ α : E, F⟮α⟯ = ⊤) : Finite (IntermediateField F E) := by diff --git a/Mathlib/FieldTheory/Separable.lean b/Mathlib/FieldTheory/Separable.lean index 159ba206bec4e..a06c06de06f6f 100644 --- a/Mathlib/FieldTheory/Separable.lean +++ b/Mathlib/FieldTheory/Separable.lean @@ -213,6 +213,7 @@ theorem separable_prod' {ι : Sort _} {f : ι → R[X]} {s : Finset ι} : h2.1.mul (ih h1.2.2 h2.2) (IsCoprime.prod_right fun i his => h1.1.2 i his <| Ne.symm <| ne_of_mem_of_not_mem his has) +open scoped Function in -- required for scoped `on` notation theorem separable_prod {ι : Sort _} [Fintype ι] {f : ι → R[X]} (h1 : Pairwise (IsCoprime on f)) (h2 : ∀ x, (f x).Separable) : (∏ x, f x).Separable := separable_prod' (fun _x _hx _y _hy hxy => h1 hxy) fun x _hx => h2 x diff --git a/Mathlib/Geometry/Euclidean/Angle/Unoriented/Affine.lean b/Mathlib/Geometry/Euclidean/Angle/Unoriented/Affine.lean index 782d2803dcdb5..1531fd8dcb8b4 100644 --- a/Mathlib/Geometry/Euclidean/Angle/Unoriented/Affine.lean +++ b/Mathlib/Geometry/Euclidean/Angle/Unoriented/Affine.lean @@ -136,9 +136,6 @@ nonrec theorem angle_le_pi (p1 p2 p3 : P) : ∠ p1 p2 p3 ≤ π := /-- The angle ∠ABA at a point is `0`, unless `A = B`. -/ theorem angle_self_of_ne (h : p ≠ p₀) : ∠ p p₀ p = 0 := angle_self <| vsub_ne_zero.2 h -@[deprecated (since := "2024-02-14")] alias angle_eq_left := angle_self_left -@[deprecated (since := "2024-02-14")] alias angle_eq_right := angle_self_right -@[deprecated (since := "2024-02-14")] alias angle_eq_of_ne := angle_self_of_ne /-- If the angle ∠ABC at a point is π, the angle ∠BAC is 0. -/ theorem angle_eq_zero_of_angle_eq_pi_left {p1 p2 p3 : P} (h : ∠ p1 p2 p3 = π) : ∠ p2 p1 p3 = 0 := by diff --git a/Mathlib/Geometry/Manifold/ContMDiffMFDeriv.lean b/Mathlib/Geometry/Manifold/ContMDiffMFDeriv.lean index 38b8fd8f39f73..e8a0505d2e9d8 100644 --- a/Mathlib/Geometry/Manifold/ContMDiffMFDeriv.lean +++ b/Mathlib/Geometry/Manifold/ContMDiffMFDeriv.lean @@ -118,7 +118,7 @@ protected theorem ContMDiffWithinAt.mfderivWithin {x₀ : N} {f : N → M → M' · apply inter_subset_left.trans rw [preimage_subset_iff] intro a ha - refine ⟨PartialEquiv.map_source _ (inter_subset_right ha : _), ?_⟩ + refine ⟨PartialEquiv.map_source _ (inter_subset_right ha :), ?_⟩ rw [mem_preimage, PartialEquiv.left_inv (extChartAt I (g x₀))] · exact hu (inter_subset_left ha) · exact (inter_subset_right ha :) diff --git a/Mathlib/Geometry/Manifold/LocalInvariantProperties.lean b/Mathlib/Geometry/Manifold/LocalInvariantProperties.lean index 2e0b7b1b65cce..704caf6b4ef1e 100644 --- a/Mathlib/Geometry/Manifold/LocalInvariantProperties.lean +++ b/Mathlib/Geometry/Manifold/LocalInvariantProperties.lean @@ -97,7 +97,7 @@ theorem congr_nhdsWithin' {s : Set H} {x : H} {f g : H → H'} (h1 : f =ᶠ[𝓝 (hG.congr_iff_nhdsWithin h1 h2).mpr hP theorem congr_iff {s : Set H} {x : H} {f g : H → H'} (h : f =ᶠ[𝓝 x] g) : P f s x ↔ P g s x := - hG.congr_iff_nhdsWithin (mem_nhdsWithin_of_mem_nhds h) (mem_of_mem_nhds h : _) + hG.congr_iff_nhdsWithin (mem_nhdsWithin_of_mem_nhds h) (mem_of_mem_nhds h :) theorem congr {s : Set H} {x : H} {f g : H → H'} (h : f =ᶠ[𝓝 x] g) (hP : P f s x) : P g s x := (hG.congr_iff h).mp hP @@ -379,7 +379,7 @@ theorem liftPropWithinAt_congr_of_eventuallyEq (h : LiftPropWithinAt P g s x) (h theorem liftPropWithinAt_congr_of_eventuallyEq_of_mem (h : LiftPropWithinAt P g s x) (h₁ : g' =ᶠ[𝓝[s] x] g) (h₂ : x ∈ s) : LiftPropWithinAt P g' s x := - liftPropWithinAt_congr_of_eventuallyEq hG h h₁ (mem_of_mem_nhdsWithin h₂ h₁ : _) + liftPropWithinAt_congr_of_eventuallyEq hG h h₁ (mem_of_mem_nhdsWithin h₂ h₁ :) theorem liftPropWithinAt_congr_iff_of_eventuallyEq (h₁ : g' =ᶠ[𝓝[s] x] g) (hx : g' x = g x) : LiftPropWithinAt P g' s x ↔ LiftPropWithinAt P g s x := diff --git a/Mathlib/Geometry/Manifold/MFDeriv/Atlas.lean b/Mathlib/Geometry/Manifold/MFDeriv/Atlas.lean index 9386f8ad261ec..92e871ba9dfed 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/Atlas.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/Atlas.lean @@ -227,11 +227,11 @@ section extChartAt variable [SmoothManifoldWithCorners I M] {s : Set M} {x y : M} {z : E} theorem hasMFDerivAt_extChartAt (h : y ∈ (chartAt H x).source) : - HasMFDerivAt I 𝓘(𝕜, E) (extChartAt I x) y (mfderiv I I (chartAt H x) y : _) := + HasMFDerivAt I 𝓘(𝕜, E) (extChartAt I x) y (mfderiv I I (chartAt H x) y :) := I.hasMFDerivAt.comp y ((mdifferentiable_chart x).mdifferentiableAt h).hasMFDerivAt theorem hasMFDerivWithinAt_extChartAt (h : y ∈ (chartAt H x).source) : - HasMFDerivWithinAt I 𝓘(𝕜, E) (extChartAt I x) s y (mfderiv I I (chartAt H x) y : _) := + HasMFDerivWithinAt I 𝓘(𝕜, E) (extChartAt I x) s y (mfderiv I I (chartAt H x) y :) := (hasMFDerivAt_extChartAt h).hasMFDerivWithinAt theorem mdifferentiableAt_extChartAt (h : y ∈ (chartAt H x).source) : diff --git a/Mathlib/Geometry/Manifold/MFDeriv/Basic.lean b/Mathlib/Geometry/Manifold/MFDeriv/Basic.lean index bd60a56508a1f..7e2cc0458ea20 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/Basic.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/Basic.lean @@ -675,7 +675,7 @@ theorem mdifferentiableWithinAt_congr_nhds {t : Set M} (hst : 𝓝[s] x = 𝓝[t protected theorem MDifferentiableWithinAt.mfderivWithin (h : MDifferentiableWithinAt I I' f s x) : mfderivWithin I I' f s x = - fderivWithin 𝕜 (writtenInExtChartAt I I' x f : _) ((extChartAt I x).symm ⁻¹' s ∩ range I) + fderivWithin 𝕜 (writtenInExtChartAt I I' x f :) ((extChartAt I x).symm ⁻¹' s ∩ range I) ((extChartAt I x) x) := by simp only [mfderivWithin, h, if_pos] @@ -687,7 +687,7 @@ theorem MDifferentiableAt.hasMFDerivAt (h : MDifferentiableAt I I' f x) : protected theorem MDifferentiableAt.mfderiv (h : MDifferentiableAt I I' f x) : mfderiv I I' f x = - fderivWithin 𝕜 (writtenInExtChartAt I I' x f : _) (range I) ((extChartAt I x) x) := by + fderivWithin 𝕜 (writtenInExtChartAt I I' x f :) (range I) ((extChartAt I x) x) := by simp only [mfderiv, h, if_pos] protected theorem HasMFDerivAt.mfderiv (h : HasMFDerivAt I I' f x f') : mfderiv I I' f x = f' := @@ -932,7 +932,7 @@ theorem HasMFDerivWithinAt.congr_mono (h : HasMFDerivWithinAt I I' f s x f') theorem HasMFDerivAt.congr_of_eventuallyEq (h : HasMFDerivAt I I' f x f') (h₁ : f₁ =ᶠ[𝓝 x] f) : HasMFDerivAt I I' f₁ x f' := by rw [← hasMFDerivWithinAt_univ] at h ⊢ - apply h.congr_of_eventuallyEq _ (mem_of_mem_nhds h₁ : _) + apply h.congr_of_eventuallyEq _ (mem_of_mem_nhds h₁ :) rwa [nhdsWithin_univ] theorem MDifferentiableWithinAt.congr_of_eventuallyEq (h : MDifferentiableWithinAt I I' f s x) @@ -1013,7 +1013,7 @@ theorem tangentMapWithin_congr (h : ∀ x ∈ s, f x = f₁ x) (p : TangentBundl theorem Filter.EventuallyEq.mfderiv_eq (hL : f₁ =ᶠ[𝓝 x] f) : mfderiv I I' f₁ x = mfderiv I I' f x := by - have A : f₁ x = f x := (mem_of_mem_nhds hL : _) + have A : f₁ x = f x := (mem_of_mem_nhds hL :) rw [← mfderivWithin_univ, ← mfderivWithin_univ] rw [← nhdsWithin_univ] at hL exact hL.mfderivWithin_eq (uniqueMDiffWithinAt_univ I) A diff --git a/Mathlib/Geometry/Manifold/MFDeriv/Defs.lean b/Mathlib/Geometry/Manifold/MFDeriv/Defs.lean index b7c3206ec3989..1629d372ecce9 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/Defs.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/Defs.lean @@ -340,7 +340,7 @@ variable (I I') in `f x`. -/ def mfderiv (f : M → M') (x : M) : TangentSpace I x →L[𝕜] TangentSpace I' (f x) := if MDifferentiableAt I I' f x then - (fderivWithin 𝕜 (writtenInExtChartAt I I' x f : E → E') (range I) ((extChartAt I x) x) : _) + (fderivWithin 𝕜 (writtenInExtChartAt I I' x f : E → E') (range I) ((extChartAt I x) x) :) else 0 variable (I I') in diff --git a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean index d3b381f0c0511..4bf23adb82cdd 100644 --- a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean +++ b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean @@ -303,7 +303,6 @@ alias closedEmbedding := isClosedEmbedding theorem isClosed_range : IsClosed (range I) := I.isClosedEmbedding.isClosed_range -@[deprecated (since := "2024-03-17")] alias closed_range := isClosed_range theorem range_eq_closure_interior : range I = closure (interior (range I)) := Subset.antisymm I.range_subset_closure_interior I.isClosed_range.closure_interior_subset diff --git a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean index edc1b664e2ba8..a3b61da96f50c 100644 --- a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean +++ b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean @@ -74,7 +74,7 @@ noncomputable def coproduct : LocallyRingedSpace where (F.obj i).isLocalRing _ exact (asIso ((colimit.ι (C := SheafedSpace.{u+1, u, u} CommRingCatMax.{u, u}) - (F ⋙ forgetToSheafedSpace) i : _).stalkMap y)).symm.commRingCatIsoToRingEquiv.isLocalRing + (F ⋙ forgetToSheafedSpace) i :).stalkMap y)).symm.commRingCatIsoToRingEquiv.isLocalRing /-- The explicit coproduct cofan for `F : discrete ι ⥤ LocallyRingedSpace`. -/ noncomputable def coproductCofan : Cocone F where @@ -82,7 +82,7 @@ noncomputable def coproductCofan : Cocone F where ι := { app := fun j => ⟨colimit.ι (C := SheafedSpace.{u+1, u, u} CommRingCatMax.{u, u}) (F ⋙ forgetToSheafedSpace) j, inferInstance⟩ - naturality := fun ⟨j⟩ ⟨j'⟩ ⟨⟨(f : j = j')⟩⟩ => by subst f; aesop } + naturality := fun ⟨j⟩ ⟨j'⟩ ⟨⟨(f : j = j')⟩⟩ => by subst f; simp } /-- The explicit coproduct cofan constructed in `coproduct_cofan` is indeed a colimit. -/ noncomputable def coproductCofanIsColimit : IsColimit (coproductCofan F) where @@ -100,7 +100,7 @@ noncomputable def coproductCofanIsColimit : IsColimit (coproductCofan F) where erw [← this, PresheafedSpace.stalkMap.congr_hom _ _ (colimit.ι_desc (C := SheafedSpace.{u+1, u, u} CommRingCatMax.{u, u}) - (forgetToSheafedSpace.mapCocone s) i : _)] + (forgetToSheafedSpace.mapCocone s) i :)] haveI : IsLocalHom (((forgetToSheafedSpace.mapCocone s).ι.app i).stalkMap y).hom := @@ -134,7 +134,7 @@ namespace HasCoequalizer @[instance] theorem coequalizer_π_app_isLocalHom (U : TopologicalSpace.Opens (coequalizer f.toShHom g.toShHom).carrier) : - IsLocalHom ((coequalizer.π f.toShHom g.toShHom : _).c.app (op U)).hom := by + IsLocalHom ((coequalizer.π f.toShHom g.toShHom :).c.app (op U)).hom := by have := ι_comp_coequalizerComparison f.toShHom g.toShHom SheafedSpace.forgetToPresheafedSpace rw [← PreservesCoequalizer.iso_hom] at this erw [SheafedSpace.congr_app this.symm (op U)] @@ -223,7 +223,7 @@ theorem imageBasicOpen_image_open : @[instance] theorem coequalizer_π_stalk_isLocalHom (x : Y) : - IsLocalHom ((coequalizer.π f.toShHom g.toShHom : _).stalkMap x).hom := by + IsLocalHom ((coequalizer.π f.toShHom g.toShHom :).stalkMap x).hom := by constructor rintro a ha rcases TopCat.Presheaf.germ_exist (C := CommRingCat) _ _ a with ⟨U, hU, s, rfl⟩ @@ -240,13 +240,13 @@ theorem coequalizer_π_stalk_isLocalHom (x : Y) : SetLike.ext' hV.symm have V_open : IsOpen ((coequalizer.π f.toShHom g.toShHom).base '' V.1) := imageBasicOpen_image_open f g U s - have VleU : (⟨(coequalizer.π f.toShHom g.toShHom).base '' V.1, V_open⟩ : _) ≤ U := + have VleU : ⟨(coequalizer.π f.toShHom g.toShHom).base '' V.1, V_open⟩ ≤ U := Set.image_subset_iff.mpr (Y.toRingedSpace.basicOpen_le _) have hxV : x ∈ V := ⟨hU, ha⟩ rw [← CommRingCat.germ_res_apply (coequalizer f.toShHom g.toShHom).presheaf (homOfLE VleU) _ (@Set.mem_image_of_mem _ _ (coequalizer.π f.toShHom g.toShHom).base x V.1 hxV) s] apply RingHom.isUnit_map - rw [← isUnit_map_iff ((coequalizer.π f.toShHom g.toShHom : _).c.app _).hom, + rw [← isUnit_map_iff ((coequalizer.π f.toShHom g.toShHom :).c.app _).hom, ← CommRingCat.comp_apply, NatTrans.naturality, CommRingCat.comp_apply, ← isUnit_map_iff (Y.presheaf.map (eqToHom hV').op).hom] -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11224): change `rw` to `erw` @@ -266,7 +266,7 @@ noncomputable def coequalizer : LocallyRingedSpace where isLocalRing x := by obtain ⟨y, rfl⟩ := (TopCat.epi_iff_surjective (coequalizer.π f.toShHom g.toShHom).base).mp inferInstance x - exact ((coequalizer.π f.toShHom g.toShHom : _).stalkMap y).hom.domain_isLocalRing + exact ((coequalizer.π f.toShHom g.toShHom :).stalkMap y).hom.domain_isLocalRing /-- The explicit coequalizer cofork of locally ringed spaces. -/ noncomputable def coequalizerCofork : Cofork f g := diff --git a/Mathlib/Geometry/RingedSpace/OpenImmersion.lean b/Mathlib/Geometry/RingedSpace/OpenImmersion.lean index 9456c6bb4f9f8..844ba7dddb0bc 100644 --- a/Mathlib/Geometry/RingedSpace/OpenImmersion.lean +++ b/Mathlib/Geometry/RingedSpace/OpenImmersion.lean @@ -130,7 +130,7 @@ theorem isoRestrict_hom_ofRestrict : (isoRestrict f).hom ≫ Y.ofRestrict _ = f erw [Category.comp_id, comp_c_app, f.c.naturality_assoc, ← X.presheaf.map_comp] trans f.c.app x ≫ X.presheaf.map (𝟙 _) · congr 1 - · erw [X.presheaf.map_id, Category.comp_id] + · simp @[reassoc (attr := simp)] theorem isoRestrict_inv_ofRestrict : (isoRestrict f).inv ≫ f = Y.ofRestrict _ := by @@ -417,7 +417,7 @@ theorem pullbackConeOfLeftLift_snd : erw [← s.pt.presheaf.map_comp, ← s.pt.presheaf.map_comp] trans s.snd.c.app x ≫ s.pt.presheaf.map (𝟙 _) · congr 1 - · rw [s.pt.presheaf.map_id]; erw [Category.comp_id] + · simp instance pullbackConeSndIsOpenImmersion : IsOpenImmersion (pullbackConeOfLeft f g).snd := by erw [CategoryTheory.Limits.PullbackCone.mk_snd] diff --git a/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean b/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean index 77f139367118c..b689c08051fcb 100644 --- a/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean +++ b/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean @@ -354,10 +354,10 @@ def ιInvApp {i : D.J} (U : Opens (D.U i).carrier) : induction Y using Opposite.rec' with | h Y => ?_ let f : Y ⟶ X := f'.unop; have : f' = f.op := rfl; clear_value f; subst this rcases f with (_ | ⟨j, k⟩ | ⟨j, k⟩) - · erw [Category.id_comp, CategoryTheory.Functor.map_id] - rw [Category.comp_id] - · erw [Category.id_comp]; congr 1 - erw [Category.id_comp] + · simp + · simp only [Functor.const_obj_obj, Functor.const_obj_map, Category.id_comp] + congr 1 + simp only [Functor.const_obj_obj, Functor.const_obj_map, Category.id_comp] -- It remains to show that the blue is equal to red + green in the original diagram. -- The proof strategy is illustrated in ![this diagram](https://i.imgur.com/mBzV1Rx.png) -- where we prove red = pink = light-blue = green = blue. diff --git a/Mathlib/Geometry/RingedSpace/PresheafedSpace/HasColimits.lean b/Mathlib/Geometry/RingedSpace/PresheafedSpace/HasColimits.lean index 4d7c3d0c180bc..1f4e389113201 100644 --- a/Mathlib/Geometry/RingedSpace/PresheafedSpace/HasColimits.lean +++ b/Mathlib/Geometry/RingedSpace/PresheafedSpace/HasColimits.lean @@ -333,8 +333,7 @@ theorem colimitPresheafObjIsoComponentwiseLimit_inv_ι_app (F : J ⥤ Presheafed suffices f_eq : f = 𝟙 _ by rw [f_eq, comp_id] erw [← (F.obj j).presheaf.map_id] change (F.obj j).presheaf.map _ ≫ _ = _ - erw [← (F.obj j).presheaf.map_comp, ← (F.obj j).presheaf.map_comp] - congr 1 + simp @[simp] theorem colimitPresheafObjIsoComponentwiseLimit_hom_π (F : J ⥤ PresheafedSpace.{_, _, v} C) diff --git a/Mathlib/Geometry/RingedSpace/Stalks.lean b/Mathlib/Geometry/RingedSpace/Stalks.lean index 3f2955405ac1d..96927f0280b51 100644 --- a/Mathlib/Geometry/RingedSpace/Stalks.lean +++ b/Mathlib/Geometry/RingedSpace/Stalks.lean @@ -169,7 +169,7 @@ instance isIso {X Y : PresheafedSpace.{_, _, v} C} (α : X ⟶ Y) [IsIso α] (x -- `X.stalk x ⟶ X.stalk ((α ≫ β).base x)`. refine ⟨eqToHom (show X.presheaf.stalk x = X.presheaf.stalk ((α ≫ β).base x) by rw [h_eq]) ≫ - (β.stalkMap (α.base x) : _), + (β.stalkMap (α.base x) :), ?_, ?_⟩ · rw [← Category.assoc, congr_point α x ((α ≫ β).base x) h_eq.symm, Category.assoc] erw [← stalkMap.comp β α (α.base x)] diff --git a/Mathlib/GroupTheory/CoprodI.lean b/Mathlib/GroupTheory/CoprodI.lean index e5bc86accb345..dbb69a5fef864 100644 --- a/Mathlib/GroupTheory/CoprodI.lean +++ b/Mathlib/GroupTheory/CoprodI.lean @@ -812,6 +812,8 @@ open Pointwise open Cardinal +open scoped Function -- required for scoped `on` notation + variable {G : Type*} [Group G] variable {H : ι → Type*} [∀ i, Group (H i)] variable (f : ∀ i, H i →* G) @@ -966,6 +968,8 @@ section PingPongLemma open Pointwise Cardinal +open scoped Function -- required for scoped `on` notation + variable [Nontrivial ι] variable {G : Type u_1} [Group G] (a : ι → G) diff --git a/Mathlib/GroupTheory/GroupAction/Basic.lean b/Mathlib/GroupTheory/GroupAction/Basic.lean index 552dee07b3bcd..520ead2860814 100644 --- a/Mathlib/GroupTheory/GroupAction/Basic.lean +++ b/Mathlib/GroupTheory/GroupAction/Basic.lean @@ -17,6 +17,13 @@ This file primarily concerns itself with orbits, stabilizers, and other objects actions. Despite this file being called `basic`, low-level helper lemmas for algebraic manipulation of `•` belong elsewhere. +## Main definitions + +* `MulAction.orbit` +* `MulAction.fixedPoints` +* `MulAction.fixedBy` +* `MulAction.stabilizer` + -/ @@ -266,6 +273,8 @@ namespace MulAction variable {G : Type*} [Group G] {α : Type*} [MulAction G α] /-- To prove inclusion of a *subgroup* in a stabilizer, it is enough to prove inclusions.-/ +@[to_additive + "To prove inclusion of a *subgroup* in a stabilizer, it is enough to prove inclusions."] theorem le_stabilizer_iff_smul_le (s : Set α) (H : Subgroup G) : H ≤ stabilizer G s ↔ ∀ g ∈ H, g • s ⊆ s := by constructor diff --git a/Mathlib/GroupTheory/GroupAction/Blocks.lean b/Mathlib/GroupTheory/GroupAction/Blocks.lean index bb890e738c69d..5655e4bb9cc96 100644 --- a/Mathlib/GroupTheory/GroupAction/Blocks.lean +++ b/Mathlib/GroupTheory/GroupAction/Blocks.lean @@ -16,6 +16,8 @@ Given `SMul G X`, an action of a type `G` on a type `X`, we define - the predicate `IsBlock G B` states that `B : Set X` is a block, which means that the sets `g • B`, for `g ∈ G`, are equal or disjoint. + Under `Group G` and `MulAction G X`, this is equivalent to the classical + definition `MulAction.IsBlock.def_one` - a bunch of lemmas that give examples of “trivial” blocks : ⊥, ⊤, singletons, and non trivial blocks: orbit of the group, orbit of a normal subgroup… @@ -25,14 +27,14 @@ The non-existence of nontrivial blocks is the definition of primitive actions. ## Results for actions on finite sets - `IsBlock.ncard_block_mul_ncard_orbit_eq` : The cardinality of a block -multiplied by the number of its translates is the cardinal of the ambient type + multiplied by the number of its translates is the cardinal of the ambient type - `IsBlock.eq_univ_of_card_lt` : a too large block is equal to `Set.univ` - `IsBlock.subsingleton_of_card_lt` : a too small block is a subsingleton - `IsBlock.of_subset` : the intersections of the translates of a finite subset -that contain a given point is a block + that contain a given point is a block ## References @@ -63,7 +65,7 @@ theorem orbit.pairwiseDisjoint : exact (orbit.eq_or_disjoint x y).resolve_right h /-- Orbits of an element form a partition -/ -@[to_additive] +@[to_additive "Orbits of an element of a partition"] theorem IsPartition.of_orbits : Setoid.IsPartition (Set.range fun a : X => orbit G a) := by apply orbit.pairwiseDisjoint.isPartition_of_exists_of_ne_empty @@ -143,8 +145,11 @@ lemma IsBlock.disjoint_smul_set_smul (hB : IsBlock G B) (hgs : ¬ g • B ⊆ s lemma IsBlock.disjoint_smul_smul_set (hB : IsBlock G B) (hgs : ¬ g • B ⊆ s • B) : Disjoint (s • B) (g • B) := (hB.disjoint_smul_set_smul hgs).symm +@[to_additive] alias ⟨IsBlock.smul_eq_smul_of_nonempty, _⟩ := isBlock_iff_smul_eq_smul_of_nonempty +@[to_additive] alias ⟨IsBlock.pairwiseDisjoint_range_smul, _⟩ := isBlock_iff_pairwiseDisjoint_range_smul +@[to_additive] alias ⟨IsBlock.smul_eq_smul_or_disjoint, _⟩ := isBlock_iff_smul_eq_smul_or_disjoint /-- A fixed block is a block. -/ @@ -277,6 +282,7 @@ protected lemma IsBlock.orbit (a : X) : IsBlock G (orbit G a) := (IsFixedBlock.o lemma isBlock_top : IsBlock (⊤ : Subgroup G) B ↔ IsBlock G B := Subgroup.topEquiv.toEquiv.forall_congr fun _ ↦ Subgroup.topEquiv.toEquiv.forall_congr_left +@[to_additive] lemma IsBlock.preimage {H Y : Type*} [Group H] [MulAction H Y] {φ : H → G} (j : Y →ₑ[φ] X) (hB : IsBlock G B) : IsBlock H (j ⁻¹' B) := by @@ -284,6 +290,7 @@ lemma IsBlock.preimage {H Y : Type*} [Group H] [MulAction H Y] rw [← Group.preimage_smul_setₛₗ, ← Group.preimage_smul_setₛₗ] at hg ⊢ exact (hB <| ne_of_apply_ne _ hg).preimage _ +@[to_additive] theorem IsBlock.image {H Y : Type*} [Group H] [MulAction H Y] {φ : G →* H} (j : X →ₑ[φ] Y) (hφ : Function.Surjective φ) (hj : Function.Injective j) @@ -292,18 +299,40 @@ theorem IsBlock.image {H Y : Type*} [Group H] [MulAction H Y] simp only [IsBlock, hφ.forall, ← image_smul_setₛₗ] exact fun g₁ g₂ hg ↦ disjoint_image_of_injective hj <| hB <| ne_of_apply_ne _ hg +@[to_additive] theorem IsBlock.subtype_val_preimage {C : SubMulAction G X} (hB : IsBlock G B) : IsBlock G (Subtype.val ⁻¹' B : Set C) := hB.preimage C.inclusion +@[to_additive] theorem isBlock_subtypeVal {C : SubMulAction G X} {B : Set C} : IsBlock G (Subtype.val '' B : Set X) ↔ IsBlock G B := by refine forall₂_congr fun g₁ g₂ ↦ ?_ rw [← SubMulAction.inclusion.coe_eq, ← image_smul_set, ← image_smul_set, ne_eq, Set.image_eq_image C.inclusion_injective, disjoint_image_iff C.inclusion_injective] +theorem _root_.AddAction.IsBlock.of_addSubgroup_of_conjugate + {G : Type*} [AddGroup G] {X : Type*} [AddAction G X] {B : Set X} + {H : AddSubgroup G} (hB : AddAction.IsBlock H B) (g : G) : + AddAction.IsBlock (H.map (AddAut.conj g).toAddMonoidHom) (g +ᵥ B) := by + rw [AddAction.isBlock_iff_vadd_eq_or_disjoint] + intro h' + obtain ⟨h, hH, hh⟩ := AddSubgroup.mem_map.mp (SetLike.coe_mem h') + simp only [AddEquiv.coe_toAddMonoidHom, AddAut.conj_apply] at hh + suffices h' +ᵥ (g +ᵥ B) = g +ᵥ (h +ᵥ B) by + simp only [this] + apply (hB.vadd_eq_or_disjoint ⟨h, hH⟩).imp + · intro hB'; congr + · exact Set.disjoint_image_of_injective (AddAction.injective g) + suffices (h' : G) +ᵥ (g +ᵥ B) = g +ᵥ (h +ᵥ B) by + exact this + rw [← hh, vadd_vadd, vadd_vadd] + erw [AddAut.conj_apply] + simp + +@[to_additive existing] theorem IsBlock.of_subgroup_of_conjugate {H : Subgroup G} (hB : IsBlock H B) (g : G) : - IsBlock (Subgroup.map (MulEquiv.toMonoidHom (MulAut.conj g)) H) (g • B) := by + IsBlock (H.map (MulAut.conj g).toMonoidHom) (g • B) := by rw [isBlock_iff_smul_eq_or_disjoint] intro h' obtain ⟨h, hH, hh⟩ := Subgroup.mem_map.mp (SetLike.coe_mem h') @@ -318,6 +347,18 @@ theorem IsBlock.of_subgroup_of_conjugate {H : Subgroup G} (hB : IsBlock H B) (g rw [← hh, smul_smul (g * h * g⁻¹) g B, smul_smul g h B, inv_mul_cancel_right] /-- A translate of a block is a block -/ +theorem _root_.AddAction.IsBlock.translate + {G : Type*} [AddGroup G] {X : Type*} [AddAction G X] (B : Set X) + (g : G) (hB : AddAction.IsBlock G B) : + AddAction.IsBlock G (g +ᵥ B) := by + rw [← AddAction.isBlock_top] at hB ⊢ + rw [← AddSubgroup.map_comap_eq_self_of_surjective (G := G) ?_ ⊤] + · apply AddAction.IsBlock.of_addSubgroup_of_conjugate + rwa [AddSubgroup.comap_top] + · exact (AddAut.conj g).surjective + +/-- A translate of a block is a block -/ +@[to_additive existing] theorem IsBlock.translate (g : G) (hB : IsBlock G B) : IsBlock G (g • B) := by rw [← isBlock_top] at hB ⊢ @@ -329,9 +370,12 @@ theorem IsBlock.translate (g : G) (hB : IsBlock G B) : variable (G) in /-- For `SMul G X`, a block system of `X` is a partition of `X` into blocks for the action of `G` -/ +@[to_additive "For `VAdd G X`, a block system of `X` is a partition of `X` into blocks + for the additive action of `G`"] def IsBlockSystem (ℬ : Set (Set X)) := Setoid.IsPartition ℬ ∧ ∀ ⦃B⦄, B ∈ ℬ → IsBlock G B -/-- Translates of a block form a block system. -/ +/-- Translates of a block form a block system -/ +@[to_additive "Translates of a block form a block system"] theorem IsBlock.isBlockSystem [hGX : MulAction.IsPretransitive G X] (hB : IsBlock G B) (hBe : B.Nonempty) : IsBlockSystem G (Set.range fun g : G => g • B) := by @@ -352,6 +396,7 @@ theorem IsBlock.isBlockSystem [hGX : MulAction.IsPretransitive G X] section Normal +@[to_additive] lemma smul_orbit_eq_orbit_smul (N : Subgroup G) [nN : N.Normal] (a : X) (g : G) : g • orbit N a = orbit N (g • a) := by simp only [orbit, Set.image_smul, Set.smul_set_range] @@ -368,6 +413,7 @@ lemma smul_orbit_eq_orbit_smul (N : Subgroup G) [nN : N.Normal] (a : X) (g : G) simp only [← mul_assoc, ← smul_smul, smul_inv_smul, inv_inv] /-- An orbit of a normal subgroup is a block -/ +@[to_additive "An orbit of a normal subgroup is a block"] theorem IsBlock.orbit_of_normal {N : Subgroup G} [N.Normal] (a : X) : IsBlock G (orbit N a) := by rw [isBlock_iff_smul_eq_or_disjoint] @@ -376,6 +422,7 @@ theorem IsBlock.orbit_of_normal {N : Subgroup G} [N.Normal] (a : X) : apply orbit.eq_or_disjoint /-- The orbits of a normal subgroup form a block system -/ +@[to_additive "The orbits of a normal subgroup form a block system"] theorem IsBlockSystem.of_normal {N : Subgroup G} [N.Normal] : IsBlockSystem G (Set.range fun a : X => orbit N a) := by constructor @@ -436,6 +483,7 @@ section Stabilizer (Wielandt, th. 7.5) -/ /-- The orbit of `a` under a subgroup containing the stabilizer of `a` is a block -/ +@[to_additive "The orbit of `a` under a subgroup containing the stabilizer of `a` is a block"] theorem IsBlock.of_orbit {H : Subgroup G} {a : X} (hH : stabilizer G a ≤ H) : IsBlock G (MulAction.orbit H a) := by rw [isBlock_iff_smul_eq_of_nonempty] @@ -447,11 +495,14 @@ theorem IsBlock.of_orbit {H : Subgroup G} {a : X} (hH : stabilizer G a ≤ H) : simpa only [mem_stabilizer_iff, InvMemClass.coe_inv, mul_smul, inv_smul_eq_iff] /-- If `B` is a block containing `a`, then the stabilizer of `B` contains the stabilizer of `a` -/ +@[to_additive + "If `B` is a block containing `a`, then the stabilizer of `B` contains the stabilizer of `a`"] theorem IsBlock.stabilizer_le (hB : IsBlock G B) {a : X} (ha : a ∈ B) : stabilizer G a ≤ stabilizer G B := fun g hg ↦ hB.smul_eq_of_nonempty ⟨a, by rwa [← hg, smul_mem_smul_set_iff], ha⟩ /-- A block containing `a` is the orbit of `a` under its stabilizer -/ +@[to_additive "A block containing `a` is the orbit of `a` under its stabilizer"] theorem IsBlock.orbit_stabilizer_eq [IsPretransitive G X] (hB : IsBlock G B) {a : X} (ha : a ∈ B) : MulAction.orbit (stabilizer G B) a = B := by ext x @@ -466,6 +517,9 @@ theorem IsBlock.orbit_stabilizer_eq [IsPretransitive G X] (hB : IsBlock G B) {a /-- A subgroup containing the stabilizer of `a` is the stabilizer of the orbit of `a` under that subgroup -/ +@[to_additive + "A subgroup containing the stabilizer of `a` + is the stabilizer of the orbit of `a` under that subgroup"] theorem stabilizer_orbit_eq {a : X} {H : Subgroup G} (hH : stabilizer G a ≤ H) : stabilizer G (orbit H a) = H := by ext g @@ -482,6 +536,9 @@ variable (G) /-- Order equivalence between blocks in `X` containing a point `a` and subgroups of `G` containing the stabilizer of `a` (Wielandt, th. 7.5)-/ +@[to_additive + "Order equivalence between blocks in `X` containing a point `a` + and subgroups of `G` containing the stabilizer of `a` (Wielandt, th. 7.5)"] def block_stabilizerOrderIso [htGX : IsPretransitive G X] (a : X) : { B : Set X // a ∈ B ∧ IsBlock G B } ≃o Set.Ici (stabilizer G a) where toFun := fun ⟨B, ha, hB⟩ => ⟨stabilizer G B, hB.stabilizer_le ha⟩ @@ -512,6 +569,7 @@ namespace IsBlock variable [IsPretransitive G X] {B : Set X} +@[to_additive] theorem ncard_block_eq_relindex (hB : IsBlock G B) {x : X} (hx : x ∈ B) : B.ncard = (stabilizer G x).relindex (stabilizer G B) := by have key : (stabilizer G x).subgroupOf (stabilizer G B) = stabilizer (stabilizer G B) x := by @@ -520,6 +578,9 @@ theorem ncard_block_eq_relindex (hB : IsBlock G B) {x : X} (hx : x ∈ B) : /-- The cardinality of the ambient space is the product of the cardinality of a block by the cardinality of the set of translates of that block -/ +@[to_additive + "The cardinality of the ambient space is the product of the cardinality of a block + by the cardinality of the set of translates of that block"] theorem ncard_block_mul_ncard_orbit_eq (hB : IsBlock G B) (hB_ne : B.Nonempty) : Set.ncard B * Set.ncard (orbit G B) = Nat.card X := by obtain ⟨x, hx⟩ := hB_ne @@ -527,11 +588,13 @@ theorem ncard_block_mul_ncard_orbit_eq (hB : IsBlock G B) (hB_ne : B.Nonempty) : Subgroup.relindex_mul_index (hB.stabilizer_le hx), index_stabilizer_of_transitive] /-- The cardinality of a block divides the cardinality of the ambient type -/ +@[to_additive "The cardinality of a block divides the cardinality of the ambient type"] theorem ncard_dvd_card (hB : IsBlock G B) (hB_ne : B.Nonempty) : Set.ncard B ∣ Nat.card X := Dvd.intro _ (hB.ncard_block_mul_ncard_orbit_eq hB_ne) /-- A too large block is equal to `univ` -/ +@[to_additive "A too large block is equal to `univ`"] theorem eq_univ_of_card_lt [hX : Finite X] (hB : IsBlock G B) (hB' : Nat.card X < Set.ncard B * 2) : B = Set.univ := by rcases Set.eq_empty_or_nonempty B with rfl | hB_ne @@ -547,6 +610,7 @@ theorem eq_univ_of_card_lt [hX : Finite X] (hB : IsBlock G B) (hB' : Nat.card X @[deprecated (since := "2024-10-29")] alias eq_univ_card_lt := eq_univ_of_card_lt /-- If a block has too many translates, then it is a (sub)singleton -/ +@[to_additive "If a block has too many translates, then it is a (sub)singleton"] theorem subsingleton_of_card_lt [Finite X] (hB : IsBlock G B) (hB' : Nat.card X < 2 * Set.ncard (orbit G B)) : B.Subsingleton := by @@ -568,7 +632,10 @@ theorem subsingleton_of_card_lt [Finite X] (hB : IsBlock G B) are just `k + ℕ`, for `k ≤ 0`, and the corresponding intersection is `ℕ`, which is not a block. (Remark by Thomas Browning) -/ /-- The intersection of the translates of a *finite* subset which contain a given point -is a block (Wielandt, th. 7.3 )-/ +is a block (Wielandt, th. 7.3)-/ +@[to_additive + "The intersection of the translates of a *finite* subset which contain a given point + is a block (Wielandt, th. 7.3)"] theorem of_subset (a : X) (hfB : B.Finite) : IsBlock G (⋂ (k : G) (_ : a ∈ k • B), k • B) := by let B' := ⋂ (k : G) (_ : a ∈ k • B), k • B diff --git a/Mathlib/GroupTheory/GroupAction/Defs.lean b/Mathlib/GroupTheory/GroupAction/Defs.lean index fa15f5e7d6253..d2cba75e644f6 100644 --- a/Mathlib/GroupTheory/GroupAction/Defs.lean +++ b/Mathlib/GroupTheory/GroupAction/Defs.lean @@ -113,7 +113,7 @@ variable (M) @[to_additive] theorem fixed_eq_iInter_fixedBy : fixedPoints M α = ⋂ m : M, fixedBy α m := Set.ext fun _ => - ⟨fun hx => Set.mem_iInter.2 fun m => hx m, fun hx m => (Set.mem_iInter.1 hx m : _)⟩ + ⟨fun hx => Set.mem_iInter.2 fun m => hx m, fun hx m => (Set.mem_iInter.1 hx m :)⟩ variable {M α} diff --git a/Mathlib/GroupTheory/GroupAction/Pointwise.lean b/Mathlib/GroupTheory/GroupAction/Pointwise.lean index 3b3918cc93570..291344f1d949b 100644 --- a/Mathlib/GroupTheory/GroupAction/Pointwise.lean +++ b/Mathlib/GroupTheory/GroupAction/Pointwise.lean @@ -29,6 +29,7 @@ import Mathlib.GroupTheory.GroupAction.Hom open Set Pointwise +@[to_additive] theorem MulAction.smul_bijective_of_is_unit {M : Type*} [Monoid M] {α : Type*} [MulAction M α] {m : M} (hm : IsUnit m) : Function.Bijective (fun (a : α) ↦ m • a) := by @@ -54,11 +55,13 @@ variable [FunLike F M N] [MulActionSemiHomClass F σ M N] -- "Left-hand side does not simplify, when using the simp lemma on itself." -- For now we will have to manually add `image_smul_setₛₗ _` to the `simp` argument list. -- TODO: when https://github.com/leanprover/lean4/issues/3107 is fixed, mark this as `@[simp]`. +@[to_additive] theorem image_smul_setₛₗ : h '' (c • s) = σ c • h '' s := by simp only [← image_smul, image_image, map_smulₛₗ h] /-- Translation of preimage is contained in preimage of translation -/ +@[to_additive] theorem smul_preimage_set_leₛₗ : c • h ⁻¹' t ⊆ h ⁻¹' (σ c • t) := by rintro x ⟨y, hy, rfl⟩ @@ -67,6 +70,7 @@ theorem smul_preimage_set_leₛₗ : variable {c} /-- General version of `preimage_smul_setₛₗ` -/ +@[to_additive] theorem preimage_smul_setₛₗ' (hc : Function.Surjective (fun (m : M) ↦ c • m)) (hc' : Function.Injective (fun (n : N) ↦ σ c • n)) : @@ -82,6 +86,7 @@ theorem preimage_smul_setₛₗ' · exact smul_preimage_set_leₛₗ M N σ h c t /-- `preimage_smul_setₛₗ` when both scalars act by unit -/ +@[to_additive] theorem preimage_smul_setₛₗ_of_units (hc : IsUnit c) (hc' : IsUnit (σ c)) : h ⁻¹' (σ c • t) = c • h ⁻¹' t := by apply preimage_smul_setₛₗ' @@ -90,6 +95,7 @@ theorem preimage_smul_setₛₗ_of_units (hc : IsUnit c) (hc' : IsUnit (σ c)) : /-- `preimage_smul_setₛₗ` in the context of a `MonoidHom` -/ +@[to_additive] theorem MonoidHom.preimage_smul_setₛₗ (σ : R →* S) {F : Type*} [FunLike F M N] [MulActionSemiHomClass F ⇑σ M N] (h : F) {c : R} (hc : IsUnit c) (t : Set N) : @@ -97,6 +103,7 @@ theorem MonoidHom.preimage_smul_setₛₗ (σ : R →* S) preimage_smul_setₛₗ_of_units M N σ h t hc (IsUnit.map σ hc) /-- `preimage_smul_setₛₗ` in the context of a `MonoidHomClass` -/ +@[to_additive] theorem preimage_smul_setₛₗ {G : Type*} [FunLike G R S] [MonoidHomClass G R S] (σ : G) {F : Type*} [FunLike F M N] [MulActionSemiHomClass F σ M N] (h : F) @@ -105,6 +112,7 @@ theorem preimage_smul_setₛₗ MonoidHom.preimage_smul_setₛₗ M N (σ : R →* S) h hc t /-- `preimage_smul_setₛₗ` in the context of a groups -/ +@[to_additive] theorem Group.preimage_smul_setₛₗ {R S : Type*} [Group R] [Group S] (σ : R → S) [MulAction R M] [MulAction S N] @@ -121,21 +129,25 @@ variable (R) variable [FunLike F M₁ M₂] [MulActionHomClass F R M₁ M₂] (c : R) (s : Set M₁) (t : Set M₂) -@[simp] -- This can be safely removed as a `@[simp]` lemma if `image_smul_setₛₗ` is readded. +-- This can be safely removed as a `@[simp]` lemma if `image_smul_setₛₗ` is readded. +@[to_additive (attr := simp)] theorem image_smul_set : h '' (c • s) = c • h '' s := image_smul_setₛₗ _ _ _ h c s +@[to_additive] theorem smul_preimage_set_le : c • h ⁻¹' t ⊆ h ⁻¹' (c • t) := smul_preimage_set_leₛₗ _ _ _ h c t variable {c} +@[to_additive] theorem preimage_smul_set (hc : IsUnit c) : h ⁻¹' (c • t) = c • h ⁻¹' t := preimage_smul_setₛₗ_of_units _ _ _ h t hc hc +@[to_additive] theorem Group.preimage_smul_set {R : Type*} [Group R] (M₁ M₂ : Type*) [MulAction R M₁] [MulAction R M₂] diff --git a/Mathlib/GroupTheory/GroupAction/SubMulAction.lean b/Mathlib/GroupTheory/GroupAction/SubMulAction.lean index babd428140d21..9478f74a0c0be 100644 --- a/Mathlib/GroupTheory/GroupAction/SubMulAction.lean +++ b/Mathlib/GroupTheory/GroupAction/SubMulAction.lean @@ -154,7 +154,15 @@ end OfTower end SetLike +/-- A SubAddAction is a set which is closed under scalar multiplication. -/ +structure SubAddAction (R : Type u) (M : Type v) [VAdd R M] : Type v where + /-- The underlying set of a `SubAddAction`. -/ + carrier : Set M + /-- The carrier set is closed under scalar multiplication. -/ + vadd_mem' : ∀ (c : R) {x : M}, x ∈ carrier → c +ᵥ x ∈ carrier + /-- A SubMulAction is a set which is closed under scalar multiplication. -/ +@[to_additive] structure SubMulAction (R : Type u) (M : Type v) [SMul R M] : Type v where /-- The underlying set of a `SubMulAction`. -/ carrier : Set M @@ -165,37 +173,44 @@ namespace SubMulAction variable [SMul R M] +@[to_additive] instance : SetLike (SubMulAction R M) M := ⟨SubMulAction.carrier, fun p q h => by cases p; cases q; congr⟩ +@[to_additive] instance : SMulMemClass (SubMulAction R M) R M where smul_mem := smul_mem' _ -@[simp] +@[to_additive (attr := simp)] theorem mem_carrier {p : SubMulAction R M} {x : M} : x ∈ p.carrier ↔ x ∈ (p : Set M) := Iff.rfl -@[ext] +@[to_additive (attr := ext)] theorem ext {p q : SubMulAction R M} (h : ∀ x, x ∈ p ↔ x ∈ q) : p = q := SetLike.ext h /-- Copy of a sub_mul_action with a new `carrier` equal to the old one. Useful to fix definitional equalities. -/ +@[to_additive "Copy of a sub_mul_action with a new `carrier` equal to the old one. + Useful to fix definitional equalities."] protected def copy (p : SubMulAction R M) (s : Set M) (hs : s = ↑p) : SubMulAction R M where carrier := s smul_mem' := hs.symm ▸ p.smul_mem' -@[simp] +@[to_additive (attr := simp)] theorem coe_copy (p : SubMulAction R M) (s : Set M) (hs : s = ↑p) : (p.copy s hs : Set M) = s := rfl +@[to_additive] theorem copy_eq (p : SubMulAction R M) (s : Set M) (hs : s = ↑p) : p.copy s hs = p := SetLike.coe_injective hs +@[to_additive] instance : Bot (SubMulAction R M) where bot := { carrier := ∅ smul_mem' := fun _c h => Set.not_mem_empty h } +@[to_additive] instance : Inhabited (SubMulAction R M) := ⟨⊥⟩ @@ -209,14 +224,16 @@ variable [SMul R M] variable (p : SubMulAction R M) variable {r : R} {x : M} +@[to_additive] theorem smul_mem (r : R) (h : x ∈ p) : r • x ∈ p := p.smul_mem' r h +@[to_additive] instance : SMul R p where smul c x := ⟨c • x.1, smul_mem _ c x.2⟩ variable {p} -@[simp, norm_cast] +@[to_additive (attr := norm_cast, simp)] theorem val_smul (r : R) (x : p) : (↑(r • x) : M) = r • (x : M) := rfl @@ -225,14 +242,16 @@ theorem val_smul (r : R) (x : p) : (↑(r • x) : M) = r • (x : M) := variable (p) /-- Embedding of a submodule `p` to the ambient space `M`. -/ +@[to_additive "Embedding of a submodule `p` to the ambient space `M`."] protected def subtype : p →[R] M where toFun := Subtype.val map_smul' := by simp [val_smul] -@[simp] +@[to_additive (attr := simp)] theorem subtype_apply (x : p) : p.subtype x = x := rfl +@[to_additive] theorem subtype_eq_val : (SubMulAction.subtype p : p → M) = Subtype.val := rfl @@ -245,14 +264,16 @@ variable [hA : SMulMemClass A R M] (S' : A) -- Prefer subclasses of `MulAction` over `SMulMemClass`. /-- A `SubMulAction` of a `MulAction` is a `MulAction`. -/ +@[to_additive "A `SubAddAction` of an `AddAction` is an `AddAction`."] instance (priority := 75) toMulAction : MulAction R S' := Subtype.coe_injective.mulAction Subtype.val (SetLike.val_smul S') /-- The natural `MulActionHom` over `R` from a `SubMulAction` of `M` to `M`. -/ +@[to_additive "The natural `AddActionHom` over `R` from a `SubAddAction` of `M` to `M`."] protected def subtype : S' →[R] M where toFun := Subtype.val; map_smul' _ _ := rfl -@[simp] +@[to_additive (attr := simp)] protected theorem coeSubtype : (SMulMemClass.subtype S' : S' → M) = Subtype.val := rfl @@ -267,28 +288,33 @@ section variable [SMul S R] [SMul S M] [IsScalarTower S R M] variable (p : SubMulAction R M) +@[to_additive] theorem smul_of_tower_mem (s : S) {x : M} (h : x ∈ p) : s • x ∈ p := by rw [← one_smul R x, ← smul_assoc] exact p.smul_mem _ h +@[to_additive] instance smul' : SMul S p where smul c x := ⟨c • x.1, smul_of_tower_mem _ c x.2⟩ +@[to_additive] instance isScalarTower : IsScalarTower S R p where smul_assoc s r x := Subtype.ext <| smul_assoc s r (x : M) +@[to_additive] instance isScalarTower' {S' : Type*} [SMul S' R] [SMul S' S] [SMul S' M] [IsScalarTower S' R M] [IsScalarTower S' S M] : IsScalarTower S' S p where smul_assoc s r x := Subtype.ext <| smul_assoc s r (x : M) -@[simp, norm_cast] +@[to_additive (attr := norm_cast, simp)] theorem val_smul_of_tower (s : S) (x : p) : ((s • x : p) : M) = s • (x : M) := rfl -@[simp] +@[to_additive (attr := simp)] theorem smul_mem_iff' {G} [Group G] [SMul G R] [MulAction G M] [IsScalarTower G R M] (g : G) {x : M} : g • x ∈ p ↔ x ∈ p := ⟨fun h => inv_smul_smul g x ▸ p.smul_of_tower_mem g⁻¹ h, p.smul_of_tower_mem g⟩ +@[to_additive] instance isCentralScalar [SMul Sᵐᵒᵖ R] [SMul Sᵐᵒᵖ M] [IsScalarTower Sᵐᵒᵖ R M] [IsCentralScalar S M] : IsCentralScalar S p where @@ -302,17 +328,20 @@ variable [Monoid S] [SMul S R] [MulAction S M] [IsScalarTower S R M] variable (p : SubMulAction R M) /-- If the scalar product forms a `MulAction`, then the subset inherits this action -/ +@[to_additive] instance mulAction' : MulAction S p where smul := (· • ·) one_smul x := Subtype.ext <| one_smul _ (x : M) mul_smul c₁ c₂ x := Subtype.ext <| mul_smul c₁ c₂ (x : M) +@[to_additive] instance mulAction : MulAction R p := p.mulAction' end /-- Orbits in a `SubMulAction` coincide with orbits in the ambient space. -/ +@[to_additive] theorem val_image_orbit {p : SubMulAction R M} (m : p) : Subtype.val '' MulAction.orbit R m = MulAction.orbit R (m : M) := (Set.range_comp _ _).symm @@ -322,15 +351,18 @@ lemma orbit_of_sub_mul {p : SubMulAction R M} (m : p) : (mul_action.orbit R m : set M) = MulAction.orbit R (m : M) := rfl -/ +@[to_additive] theorem val_preimage_orbit {p : SubMulAction R M} (m : p) : Subtype.val ⁻¹' MulAction.orbit R (m : M) = MulAction.orbit R m := by rw [← val_image_orbit, Subtype.val_injective.preimage_image] +@[to_additive] lemma mem_orbit_subMul_iff {p : SubMulAction R M} {x m : p} : x ∈ MulAction.orbit R m ↔ (x : M) ∈ MulAction.orbit R (m : M) := by rw [← val_preimage_orbit, Set.mem_preimage] /-- Stabilizers in monoid SubMulAction coincide with stabilizers in the ambient space -/ +@[to_additive] theorem stabilizer_of_subMul.submonoid {p : SubMulAction R M} (m : p) : MulAction.stabilizerSubmonoid R m = MulAction.stabilizerSubmonoid R (m : M) := by ext @@ -342,6 +374,7 @@ section MulActionGroup variable [Group R] [MulAction R M] +@[to_additive] lemma orbitRel_of_subMul (p : SubMulAction R M) : MulAction.orbitRel R p = (MulAction.orbitRel R M).comap Subtype.val := by refine Setoid.ext_iff.2 (fun x y ↦ ?_) @@ -349,6 +382,7 @@ lemma orbitRel_of_subMul (p : SubMulAction R M) : exact mem_orbit_subMul_iff /-- Stabilizers in group SubMulAction coincide with stabilizers in the ambient space -/ +@[to_additive] theorem stabilizer_of_subMul {p : SubMulAction R M} (m : p) : MulAction.stabilizer R m = MulAction.stabilizer R (m : M) := by rw [← Subgroup.toSubmonoid_inj] @@ -419,23 +453,28 @@ variable {M α : Type*} [Monoid M] [MulAction M α] /-- The inclusion of a SubMulAction into the ambient set, as an equivariant map -/ +@[to_additive "The inclusion of a SubAddAction into the ambient set, as an equivariant map."] def inclusion (s : SubMulAction M α) : s →[M] α where -- The inclusion map of the inclusion of a SubMulAction toFun := Subtype.val -- The commutation property map_smul' _ _ := rfl +@[to_additive] theorem inclusion.toFun_eq_coe (s : SubMulAction M α) : s.inclusion.toFun = Subtype.val := rfl +@[to_additive] theorem inclusion.coe_eq (s : SubMulAction M α) : ⇑s.inclusion = Subtype.val := rfl +@[to_additive] lemma image_inclusion (s : SubMulAction M α) : Set.range s.inclusion = s.carrier := by rw [inclusion.coe_eq] exact Subtype.range_coe +@[to_additive] lemma inclusion_injective (s : SubMulAction M α) : Function.Injective s.inclusion := Subtype.val_injective diff --git a/Mathlib/GroupTheory/Index.lean b/Mathlib/GroupTheory/Index.lean index 6ad32109acca6..f8e36f80728c3 100644 --- a/Mathlib/GroupTheory/Index.lean +++ b/Mathlib/GroupTheory/Index.lean @@ -617,12 +617,12 @@ namespace MulAction variable (G : Type*) {X : Type*} [Group G] [MulAction G X] (x : X) -theorem index_stabilizer : +@[to_additive] theorem index_stabilizer : (stabilizer G x).index = (orbit G x).ncard := (Nat.card_congr (MulAction.orbitEquivQuotientStabilizer G x)).symm.trans (Set.Nat.card_coe_set_eq (orbit G x)) -theorem index_stabilizer_of_transitive [IsPretransitive G X] : +@[to_additive] theorem index_stabilizer_of_transitive [IsPretransitive G X] : (stabilizer G x).index = Nat.card X := by rw [index_stabilizer, orbit_eq_univ, Set.ncard_univ] diff --git a/Mathlib/GroupTheory/OrderOfElement.lean b/Mathlib/GroupTheory/OrderOfElement.lean index a5f1e2aead8f6..514fd7841f060 100644 --- a/Mathlib/GroupTheory/OrderOfElement.lean +++ b/Mathlib/GroupTheory/OrderOfElement.lean @@ -263,9 +263,6 @@ protected lemma IsOfFinOrder.mem_powers_iff_mem_range_orderOf [DecidableEq G] protected lemma IsOfFinOrder.powers_eq_image_range_orderOf [DecidableEq G] (hx : IsOfFinOrder x) : (Submonoid.powers x : Set G) = (Finset.range (orderOf x)).image (x ^ ·) := Set.ext fun _ ↦ hx.mem_powers_iff_mem_range_orderOf -@[deprecated (since := "2024-02-21")] -alias IsOfFinAddOrder.powers_eq_image_range_orderOf := - IsOfFinAddOrder.multiples_eq_image_range_addOrderOf @[to_additive] theorem pow_eq_one_iff_modEq : x ^ n = 1 ↔ n ≡ 0 [MOD orderOf x] := by @@ -1089,19 +1086,14 @@ variable [Monoid α] [Monoid β] {x : α × β} {a : α} {b : β} @[to_additive] protected theorem Prod.orderOf (x : α × β) : orderOf x = (orderOf x.1).lcm (orderOf x.2) := minimalPeriod_prod_map _ _ _ -@[deprecated (since := "2024-02-21")] alias Prod.add_orderOf := Prod.addOrderOf @[to_additive] theorem orderOf_fst_dvd_orderOf : orderOf x.1 ∣ orderOf x := minimalPeriod_fst_dvd -@[deprecated (since := "2024-02-21")] -alias add_orderOf_fst_dvd_add_orderOf := addOrderOf_fst_dvd_addOrderOf @[to_additive] theorem orderOf_snd_dvd_orderOf : orderOf x.2 ∣ orderOf x := minimalPeriod_snd_dvd -@[deprecated (since := "2024-02-21")] alias -add_orderOf_snd_dvd_add_orderOf := addOrderOf_snd_dvd_addOrderOf @[to_additive] theorem IsOfFinOrder.fst {x : α × β} (hx : IsOfFinOrder x) : IsOfFinOrder x.1 := diff --git a/Mathlib/GroupTheory/Perm/Cycle/Basic.lean b/Mathlib/GroupTheory/Perm/Cycle/Basic.lean index 84311136f832a..7a9275a5282d0 100644 --- a/Mathlib/GroupTheory/Perm/Cycle/Basic.lean +++ b/Mathlib/GroupTheory/Perm/Cycle/Basic.lean @@ -5,7 +5,7 @@ Authors: Chris Hughes, Yaël Dillies -/ import Mathlib.Algebra.Module.BigOperators -import Mathlib.Data.Fintype.Perm +import Mathlib.Data.List.Iterate import Mathlib.GroupTheory.Perm.Finite import Mathlib.GroupTheory.Perm.List @@ -205,12 +205,13 @@ instance (priority := 100) [DecidableEq α] : DecidableRel (SameCycle (1 : Perm decidable_of_iff (x = y) sameCycle_one.symm instance [Fintype α] [DecidableEq α] (f : Perm α) : DecidableRel (SameCycle f) := fun x y => - decidable_of_iff (∃ n ∈ List.range (Fintype.card (Perm α)), (f ^ n) x = y) - ⟨fun ⟨n, _, hn⟩ => ⟨n, hn⟩, fun ⟨i, hi⟩ => ⟨(i % orderOf f).natAbs, - List.mem_range.2 (Int.ofNat_lt.1 <| by + decidable_of_iff (y ∈ List.iterate f x (Fintype.card (Perm α))) <| by + simp only [List.mem_iterate, iterate_eq_pow, eq_comm (a := y)] + exact ⟨fun ⟨n, _, hn⟩ => ⟨n, hn⟩, fun ⟨i, hi⟩ => ⟨(i % orderOf f).natAbs, + Int.ofNat_lt.1 <| by rw [Int.natAbs_of_nonneg (Int.emod_nonneg _ <| Int.natCast_ne_zero.2 (orderOf_pos _).ne')] refine (Int.emod_lt _ <| Int.natCast_ne_zero_iff_pos.2 <| orderOf_pos _).trans_le ?_ - simp [orderOf_le_card_univ]), + simp [orderOf_le_card_univ], by rw [← zpow_natCast, Int.natAbs_of_nonneg (Int.emod_nonneg _ <| Int.natCast_ne_zero_iff_pos.2 <| orderOf_pos _), zpow_mod_orderOf, hi]⟩⟩ diff --git a/Mathlib/GroupTheory/Perm/Cycle/Concrete.lean b/Mathlib/GroupTheory/Perm/Cycle/Concrete.lean index 50e4396f56fa1..4f238742343a8 100644 --- a/Mathlib/GroupTheory/Perm/Cycle/Concrete.lean +++ b/Mathlib/GroupTheory/Perm/Cycle/Concrete.lean @@ -193,7 +193,7 @@ variable [Fintype α] [DecidableEq α] (p : Equiv.Perm α) (x : α) until looping. That means when `f x = x`, `toList f x = []`. -/ def toList : List α := - (List.range (cycleOf p x).support.card).map fun k => (p ^ k) x + List.iterate p x (cycleOf p x).support.card @[simp] theorem toList_one : toList (1 : Perm α) x = [] := by simp [toList, cycleOf_one] @@ -223,7 +223,7 @@ theorem toList_get_zero (h : x ∈ p.support) : variable {p} {x} theorem mem_toList_iff {y : α} : y ∈ toList p x ↔ SameCycle p x y ∧ x ∈ p.support := by - simp only [toList, mem_range, mem_map] + simp only [toList, mem_iterate, iterate_eq_pow, eq_comm (a := y)] constructor · rintro ⟨n, hx, rfl⟩ refine ⟨⟨n, rfl⟩, ?_⟩ @@ -314,8 +314,8 @@ theorem toList_formPerm_nontrivial (l : List α) (hl : 2 ≤ l.length) (hn : Nod simp [Nat.succ_le_succ_iff] at hl rw [toList, hc.cycleOf_eq (mem_support.mp _), hs, card_toFinset, dedup_eq_self.mpr hn] · refine ext_getElem (by simp) fun k hk hk' => ?_ - simp only [get_eq_getElem, formPerm_pow_apply_getElem _ hn, zero_add, getElem_map, - getElem_range, Nat.mod_eq_of_lt hk'] + simp only [get_eq_getElem, getElem_iterate, iterate_eq_pow, formPerm_pow_apply_getElem _ hn, + zero_add, Nat.mod_eq_of_lt hk'] · simp [hs] theorem toList_formPerm_isRotated_self (l : List α) (hl : 2 ≤ l.length) (hn : Nodup l) (x : α) diff --git a/Mathlib/GroupTheory/Perm/Cycle/Factors.lean b/Mathlib/GroupTheory/Perm/Cycle/Factors.lean index 990c4810661c5..a3fe94c15f780 100644 --- a/Mathlib/GroupTheory/Perm/Cycle/Factors.lean +++ b/Mathlib/GroupTheory/Perm/Cycle/Factors.lean @@ -4,14 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Yaël Dillies -/ -import Mathlib.Algebra.Module.BigOperators -import Mathlib.Data.Finset.NoncommProd -import Mathlib.Data.Fintype.Perm -import Mathlib.Data.Int.ModEq -import Mathlib.GroupTheory.Perm.List -import Mathlib.GroupTheory.Perm.Sign -import Mathlib.Logic.Equiv.Fintype - import Mathlib.GroupTheory.Perm.Cycle.Basic /-! diff --git a/Mathlib/GroupTheory/Perm/Finite.lean b/Mathlib/GroupTheory/Perm/Finite.lean index feec477ac45f6..87b337db2fa4f 100644 --- a/Mathlib/GroupTheory/Perm/Finite.lean +++ b/Mathlib/GroupTheory/Perm/Finite.lean @@ -63,7 +63,7 @@ theorem perm_inv_on_of_perm_on_finset {s : Finset α} {f : Perm α} (h : ∀ x simp only [inv_apply_self] theorem perm_inv_mapsTo_of_mapsTo (f : Perm α) {s : Set α} [Finite s] (h : Set.MapsTo f s s) : - Set.MapsTo (f⁻¹ : _) s s := by + Set.MapsTo (f⁻¹ :) s s := by cases nonempty_fintype s exact fun x hx => Set.mem_toFinset.mp <| @@ -73,7 +73,7 @@ theorem perm_inv_mapsTo_of_mapsTo (f : Perm α) {s : Set α} [Finite s] (h : Set @[simp] theorem perm_inv_mapsTo_iff_mapsTo {f : Perm α} {s : Set α} [Finite s] : - Set.MapsTo (f⁻¹ : _) s s ↔ Set.MapsTo f s s := + Set.MapsTo (f⁻¹ :) s s ↔ Set.MapsTo f s s := ⟨perm_inv_mapsTo_of_mapsTo f⁻¹, perm_inv_mapsTo_of_mapsTo f⟩ theorem perm_inv_on_of_perm_on_finite {f : Perm α} {p : α → Prop} [Finite { x // p x }] diff --git a/Mathlib/GroupTheory/QuotientGroup/Basic.lean b/Mathlib/GroupTheory/QuotientGroup/Basic.lean index 58622e2d53e83..ac753e21b448b 100644 --- a/Mathlib/GroupTheory/QuotientGroup/Basic.lean +++ b/Mathlib/GroupTheory/QuotientGroup/Basic.lean @@ -27,6 +27,9 @@ proves Noether's first and second isomorphism theorems. `N` of a group `G`. * `QuotientGroup.quotientQuotientEquivQuotient`: Noether's third isomorphism theorem, the canonical isomorphism between `(G / N) / (M / N)` and `G / M`, where `N ≤ M`. +* `QuotientGroup.comapMk'OrderIso`: The correspondence theorem, a lattice + isomorphism between the lattice of subgroups of `G ⧸ N` and the sublattice + of subgroups of `G` containing `N`. ## Tags @@ -321,6 +324,35 @@ def quotientQuotientEquivQuotient : (G ⧸ N) ⧸ M.map (QuotientGroup.mk' N) end ThirdIsoThm +section CorrespTheorem + +-- All these theorems are primed because `QuotientGroup.mk'` is. +set_option linter.docPrime false + +@[to_additive] +theorem le_comap_mk' (N : Subgroup G) [N.Normal] (H : Subgroup (G ⧸ N)) : + N ≤ Subgroup.comap (QuotientGroup.mk' N) H := by + simpa using Subgroup.comap_mono (f := mk' N) bot_le + +@[to_additive (attr := simp)] +theorem comap_map_mk' (N H : Subgroup G) [N.Normal] : + Subgroup.comap (mk' N) (Subgroup.map (mk' N) H) = N ⊔ H := by + simp [Subgroup.comap_map_eq, sup_comm] + +/-- The **correspondence theorem**, or lattice theorem, + or fourth isomorphism theorem for multiplicative groups-/ +@[to_additive "The **correspondence theorem**, or lattice theorem, + or fourth isomorphism theorem for additive groups"] +def comapMk'OrderIso (N : Subgroup G) [hn : N.Normal] : + Subgroup (G ⧸ N) ≃o { H : Subgroup G // N ≤ H } where + toFun H' := ⟨Subgroup.comap (mk' N) H', le_comap_mk' N _⟩ + invFun H := Subgroup.map (mk' N) H + left_inv H' := Subgroup.map_comap_eq_self <| by simp + right_inv := fun ⟨H, hH⟩ => Subtype.ext_val <| by simpa + map_rel_iff' := Subgroup.comap_le_comap_of_surjective <| mk'_surjective _ + +end CorrespTheorem + section trivial @[to_additive] diff --git a/Mathlib/GroupTheory/QuotientGroup/Defs.lean b/Mathlib/GroupTheory/QuotientGroup/Defs.lean index d33a94cc017c6..a4bed8c014303 100644 --- a/Mathlib/GroupTheory/QuotientGroup/Defs.lean +++ b/Mathlib/GroupTheory/QuotientGroup/Defs.lean @@ -84,13 +84,20 @@ See note [partially-applied ext lemmas]. -/ See note [partially-applied ext lemmas]. "] theorem monoidHom_ext ⦃f g : G ⧸ N →* M⦄ (h : f.comp (mk' N) = g.comp (mk' N)) : f = g := - MonoidHom.ext fun x => QuotientGroup.induction_on x <| (DFunLike.congr_fun h : _) + MonoidHom.ext fun x => QuotientGroup.induction_on x <| (DFunLike.congr_fun h :) @[to_additive (attr := simp)] theorem eq_one_iff {N : Subgroup G} [N.Normal] (x : G) : (x : G ⧸ N) = 1 ↔ x ∈ N := by refine QuotientGroup.eq.trans ?_ rw [mul_one, Subgroup.inv_mem_iff] +/- Note: `range_mk'` is a lemma about the primed constructor `QuotientGroup.mk'`, not a + modified version of some `range_mk`. -/ +set_option linter.docPrime false in +@[to_additive (attr := simp)] +theorem range_mk' : (QuotientGroup.mk' N).range = ⊤ := + MonoidHom.range_eq_top.mpr (mk'_surjective N) + @[to_additive] theorem ker_le_range_iff {I : Type w} [Group I] (f : G →* H) [f.range.Normal] (g : H →* I) : g.ker ≤ f.range ↔ (mk' f.range).comp g.ker.subtype = 1 := diff --git a/Mathlib/GroupTheory/SpecificGroups/Alternating.lean b/Mathlib/GroupTheory/SpecificGroups/Alternating.lean index df3d8243d1cdb..f24b630fe1377 100644 --- a/Mathlib/GroupTheory/SpecificGroups/Alternating.lean +++ b/Mathlib/GroupTheory/SpecificGroups/Alternating.lean @@ -231,7 +231,7 @@ theorem normalClosure_finRotate_five : normalClosure ({⟨finRotate 5, SetLike.mem_coe.1 (subset_normalClosure (Set.mem_singleton _)) exact (mul_mem (Subgroup.normalClosure_normal.conj_mem _ h -- Porting note: added `: _` - ⟨Fin.cycleRange 2, Fin.isThreeCycle_cycleRange_two.mem_alternatingGroup⟩) (inv_mem h) : _)) + ⟨Fin.cycleRange 2, Fin.isThreeCycle_cycleRange_two.mem_alternatingGroup⟩) (inv_mem h) :)) /-- The normal closure of $(04)(13)$ within $A_5$ is the whole group. This will be used to show that the normal closure of any permutation of cycle type $(2,2)$ is the whole group. diff --git a/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean b/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean index 5ea4adad723cf..b469bc7d71d9f 100644 --- a/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean +++ b/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean @@ -105,8 +105,6 @@ theorem MonoidHom.map_cyclic [h : IsCyclic G] (σ : G →* G) : refine ⟨m, fun g => ?_⟩ obtain ⟨n, rfl⟩ := hG g rw [MonoidHom.map_zpow, ← hm, ← zpow_mul, ← zpow_mul'] -@[deprecated (since := "2024-02-21")] alias -MonoidAddHom.map_add_cyclic := AddMonoidHom.map_addCyclic @[to_additive] lemma isCyclic_iff_exists_orderOf_eq_natCard [Finite α] : @@ -132,9 +130,6 @@ theorem isCyclic_of_orderOf_eq_card [Finite α] (x : α) (hx : orderOf x = Nat.c IsCyclic α := isCyclic_iff_exists_orderOf_eq_natCard.mpr ⟨x, hx⟩ -@[deprecated (since := "2024-02-21")] -alias isAddCyclic_of_orderOf_eq_card := isAddCyclic_of_addOrderOf_eq_card - @[to_additive] theorem Subgroup.eq_bot_or_eq_top_of_prime_card (H : Subgroup G) [hp : Fact (Nat.card G).Prime] : H = ⊥ ∨ H = ⊤ := by @@ -324,8 +319,6 @@ theorem IsCyclic.card_pow_eq_one_le [DecidableEq α] [Fintype α] [IsCyclic α] rw [Nat.mul_div_cancel_left _ (gcd_pos_of_pos_left _ hn0), gcd_mul_left_left, hm, Nat.mul_div_cancel _ hm0] exact le_of_dvd hn0 (Nat.gcd_dvd_left _ _) -@[deprecated (since := "2024-02-21")] -alias IsAddCyclic.card_pow_eq_one_le := IsAddCyclic.card_nsmul_eq_zero_le end Classical @@ -478,9 +471,6 @@ theorem isCyclic_of_card_pow_eq_one_le : IsCyclic α := let ⟨x, hx⟩ := this isCyclic_of_orderOf_eq_card x (Finset.mem_filter.1 hx).2 -@[deprecated (since := "2024-02-21")] -alias isAddCyclic_of_card_pow_eq_one_le := isAddCyclic_of_card_nsmul_eq_zero_le - end Totient @[to_additive] @@ -488,9 +478,6 @@ lemma IsCyclic.card_orderOf_eq_totient [IsCyclic α] [Fintype α] {d : ℕ} (hd #{a : α | orderOf a = d} = totient d := by classical apply card_orderOf_eq_totient_aux₂ (fun n => IsCyclic.card_pow_eq_one_le) hd -@[deprecated (since := "2024-02-21")] -alias IsAddCyclic.card_orderOf_eq_totient := IsAddCyclic.card_addOrderOf_eq_totient - /-- A finite group of prime order is simple. -/ @[to_additive "A finite group of prime order is simple."] theorem isSimpleGroup_of_prime_card {p : ℕ} [hp : Fact p.Prime] @@ -532,9 +519,6 @@ theorem commutative_of_cyclic_center_quotient [IsCyclic G'] (f : G →* G') (hf _ = y ^ m * y ^ n * y ^ (-m) * (y ^ (-n) * b * a) := by rw [mem_center_iff.1 hb] _ = b * a := by group -@[deprecated (since := "2024-02-21")] -alias commutative_of_add_cyclic_center_quotient := commutative_of_addCyclic_center_quotient - /-- A group is commutative if the quotient by the center is cyclic. -/ @[to_additive "A group is commutative if the quotient by the center is cyclic."] diff --git a/Mathlib/GroupTheory/SpecificGroups/KleinFour.lean b/Mathlib/GroupTheory/SpecificGroups/KleinFour.lean index e5a725fd32218..2ad3ed38cd7f4 100644 --- a/Mathlib/GroupTheory/SpecificGroups/KleinFour.lean +++ b/Mathlib/GroupTheory/SpecificGroups/KleinFour.lean @@ -113,7 +113,7 @@ lemma eq_finset_univ [Fintype G] [DecidableEq G] rw [card_four'] repeat rw [card_insert_of_not_mem] on_goal 4 => simpa using mul_not_mem_of_exponent_two (by simp) hx hy hxy - all_goals aesop + all_goals simp_all @[to_additive] lemma eq_mul_of_ne_all {x y z : G} (hx : x ≠ 1) diff --git a/Mathlib/GroupTheory/Submonoid/Inverses.lean b/Mathlib/GroupTheory/Submonoid/Inverses.lean index 997f9b0ea0a44..8af3344ef8ae6 100644 --- a/Mathlib/GroupTheory/Submonoid/Inverses.lean +++ b/Mathlib/GroupTheory/Submonoid/Inverses.lean @@ -70,7 +70,7 @@ theorem leftInv_leftInv_le : S.leftInv.leftInv ≤ S := by rw [← mul_one x, ← h₁, ← mul_assoc, h₂, one_mul] @[to_additive] -theorem unit_mem_leftInv (x : Mˣ) (hx : (x : M) ∈ S) : ((x⁻¹ : _) : M) ∈ S.leftInv := +theorem unit_mem_leftInv (x : Mˣ) (hx : (x : M) ∈ S) : ((x⁻¹ :) : M) ∈ S.leftInv := ⟨⟨x, hx⟩, x.inv_val⟩ @[to_additive] diff --git a/Mathlib/LinearAlgebra/Alternating/Basic.lean b/Mathlib/LinearAlgebra/Alternating/Basic.lean index cde969096e601..f45db0f96d8da 100644 --- a/Mathlib/LinearAlgebra/Alternating/Basic.lean +++ b/Mathlib/LinearAlgebra/Alternating/Basic.lean @@ -820,11 +820,11 @@ def alternatization : MultilinearMap R (fun _ : ι => M) N' →+ M [⋀^ι]→ zero_apply, smul_zero, Finset.sum_const_zero, AlternatingMap.zero_apply] theorem alternatization_def (m : MultilinearMap R (fun _ : ι => M) N') : - ⇑(alternatization m) = (∑ σ : Perm ι, Equiv.Perm.sign σ • m.domDomCongr σ : _) := + ⇑(alternatization m) = (∑ σ : Perm ι, Equiv.Perm.sign σ • m.domDomCongr σ :) := rfl theorem alternatization_coe (m : MultilinearMap R (fun _ : ι => M) N') : - ↑(alternatization m) = (∑ σ : Perm ι, Equiv.Perm.sign σ • m.domDomCongr σ : _) := + ↑(alternatization m) = (∑ σ : Perm ι, Equiv.Perm.sign σ • m.domDomCongr σ :) := coe_injective rfl theorem alternatization_apply (m : MultilinearMap R (fun _ : ι => M) N') (v : ι → M) : diff --git a/Mathlib/LinearAlgebra/BilinearForm/Properties.lean b/Mathlib/LinearAlgebra/BilinearForm/Properties.lean index 47aa85348d3ba..1c1eeeec108bc 100644 --- a/Mathlib/LinearAlgebra/BilinearForm/Properties.lean +++ b/Mathlib/LinearAlgebra/BilinearForm/Properties.lean @@ -90,10 +90,10 @@ protected theorem eq (H : B.IsSymm) (x y : M) : B x y = B y x := theorem isRefl (H : B.IsSymm) : B.IsRefl := fun x y H1 => H x y ▸ H1 protected theorem add {B₁ B₂ : BilinForm R M} (hB₁ : B₁.IsSymm) (hB₂ : B₂.IsSymm) : - (B₁ + B₂).IsSymm := fun x y => (congr_arg₂ (· + ·) (hB₁ x y) (hB₂ x y) : _) + (B₁ + B₂).IsSymm := fun x y => (congr_arg₂ (· + ·) (hB₁ x y) (hB₂ x y) :) protected theorem sub {B₁ B₂ : BilinForm R₁ M₁} (hB₁ : B₁.IsSymm) (hB₂ : B₂.IsSymm) : - (B₁ - B₂).IsSymm := fun x y => (congr_arg₂ Sub.sub (hB₁ x y) (hB₂ x y) : _) + (B₁ - B₂).IsSymm := fun x y => (congr_arg₂ Sub.sub (hB₁ x y) (hB₂ x y) :) protected theorem neg {B : BilinForm R₁ M₁} (hB : B.IsSymm) : (-B).IsSymm := fun x y => congr_arg Neg.neg (hB x y) @@ -133,7 +133,7 @@ theorem eq_of_add_add_eq_zero [IsCancelAdd R] {a b c : M} (H : B.IsAlt) (hAdd : B a b = B b c := LinearMap.IsAlt.eq_of_add_add_eq_zero H hAdd protected theorem add {B₁ B₂ : BilinForm R M} (hB₁ : B₁.IsAlt) (hB₂ : B₂.IsAlt) : (B₁ + B₂).IsAlt := - fun x => (congr_arg₂ (· + ·) (hB₁ x) (hB₂ x) : _).trans <| add_zero _ + fun x => (congr_arg₂ (· + ·) (hB₁ x) (hB₂ x) :).trans <| add_zero _ protected theorem sub {B₁ B₂ : BilinForm R₁ M₁} (hB₁ : B₁.IsAlt) (hB₂ : B₂.IsAlt) : (B₁ - B₂).IsAlt := fun x => (congr_arg₂ Sub.sub (hB₁ x) (hB₂ x)).trans <| sub_zero _ diff --git a/Mathlib/LinearAlgebra/BilinearMap.lean b/Mathlib/LinearAlgebra/BilinearMap.lean index 176ef9be83217..e1b2be24a3275 100644 --- a/Mathlib/LinearAlgebra/BilinearMap.lean +++ b/Mathlib/LinearAlgebra/BilinearMap.lean @@ -196,7 +196,7 @@ def restrictScalars₁₂ (B : M →ₗ[R] N →ₗ[S] Pₗ) : M →ₗ[R'] N theorem restrictScalars₁₂_injective : Function.Injective (LinearMap.restrictScalars₁₂ R' S' : (M →ₗ[R] N →ₗ[S] Pₗ) → (M →ₗ[R'] N →ₗ[S'] Pₗ)) := - fun _ _ h ↦ ext₂ (congr_fun₂ h : _) + fun _ _ h ↦ ext₂ (congr_fun₂ h :) @[simp] theorem restrictScalars₁₂_inj {B B' : M →ₗ[R] N →ₗ[S] Pₗ} : diff --git a/Mathlib/LinearAlgebra/CliffordAlgebra/BaseChange.lean b/Mathlib/LinearAlgebra/CliffordAlgebra/BaseChange.lean index 5651b6f8c5e13..8bf95f6eab875 100644 --- a/Mathlib/LinearAlgebra/CliffordAlgebra/BaseChange.lean +++ b/Mathlib/LinearAlgebra/CliffordAlgebra/BaseChange.lean @@ -161,7 +161,7 @@ theorem toBaseChange_comp_ofBaseChange (Q : QuadraticForm R V) : @[simp] theorem toBaseChange_ofBaseChange (Q : QuadraticForm R V) (x : A ⊗[R] CliffordAlgebra Q) : toBaseChange A Q (ofBaseChange A Q x) = x := - AlgHom.congr_fun (toBaseChange_comp_ofBaseChange A Q : _) x + AlgHom.congr_fun (toBaseChange_comp_ofBaseChange A Q :) x theorem ofBaseChange_comp_toBaseChange (Q : QuadraticForm R V) : (ofBaseChange A Q).comp (toBaseChange A Q) = AlgHom.id _ _ := by @@ -173,7 +173,7 @@ theorem ofBaseChange_comp_toBaseChange (Q : QuadraticForm R V) : @[simp] theorem ofBaseChange_toBaseChange (Q : QuadraticForm R V) (x : CliffordAlgebra (Q.baseChange A)) : ofBaseChange A Q (toBaseChange A Q x) = x := - AlgHom.congr_fun (ofBaseChange_comp_toBaseChange A Q : _) x + AlgHom.congr_fun (ofBaseChange_comp_toBaseChange A Q :) x /-- Base-changing the vector space of a clifford algebra is isomorphic as an A-algebra to base-changing the clifford algebra itself; <|Cℓ(A ⊗_R V, Q_A) ≅ A ⊗_R Cℓ(V, Q)<|. diff --git a/Mathlib/LinearAlgebra/CliffordAlgebra/EvenEquiv.lean b/Mathlib/LinearAlgebra/CliffordAlgebra/EvenEquiv.lean index 27682d972414a..9699eda64fcc2 100644 --- a/Mathlib/LinearAlgebra/CliffordAlgebra/EvenEquiv.lean +++ b/Mathlib/LinearAlgebra/CliffordAlgebra/EvenEquiv.lean @@ -228,7 +228,7 @@ def evenToNeg (Q' : QuadraticForm R M) (h : Q' = -Q) : -- Porting note: added `letI`s letI : AddCommGroup (even Q') := AddSubgroupClass.toAddCommGroup _ letI : HasDistribNeg (even Q') := NonUnitalNonAssocRing.toHasDistribNeg - { bilin := -(even.ι Q' : _).bilin + { bilin := -(even.ι Q' :).bilin contract := fun m => by simp_rw [LinearMap.neg_apply, EvenHom.contract, h, QuadraticMap.neg_apply, map_neg, neg_neg] contract_mid := fun m₁ m₂ m₃ => by @@ -244,11 +244,7 @@ theorem evenToNeg_ι (Q' : QuadraticForm R M) (h : Q' = -Q) (m₁ m₂ : M) : theorem evenToNeg_comp_evenToNeg (Q' : QuadraticForm R M) (h : Q' = -Q) (h' : Q = -Q') : (evenToNeg Q' Q h').comp (evenToNeg Q Q' h) = AlgHom.id R _ := by ext m₁ m₂ : 4 - dsimp only [EvenHom.compr₂_bilin, LinearMap.compr₂_apply, AlgHom.toLinearMap_apply, - AlgHom.comp_apply, AlgHom.id_apply] - rw [evenToNeg_ι] - -- Needed to use `RingHom.map_neg` to avoid a timeout and now `erw` https://github.com/leanprover-community/mathlib4/pull/8386 - erw [RingHom.map_neg, evenToNeg_ι, neg_neg] + simp [evenToNeg_ι] /-- The even subalgebras of the algebras with quadratic form `Q` and `-Q` are isomorphic. diff --git a/Mathlib/LinearAlgebra/Dimension/Basic.lean b/Mathlib/LinearAlgebra/Dimension/Basic.lean index 5e63e9aa3eb1a..124960fd5f2ea 100644 --- a/Mathlib/LinearAlgebra/Dimension/Basic.lean +++ b/Mathlib/LinearAlgebra/Dimension/Basic.lean @@ -68,7 +68,7 @@ end namespace LinearIndependent -variable [Ring R] [AddCommGroup M] [Module R M] +variable [Semiring R] [AddCommMonoid M] [Module R M] variable [Nontrivial R] diff --git a/Mathlib/LinearAlgebra/Dimension/Constructions.lean b/Mathlib/LinearAlgebra/Dimension/Constructions.lean index 4caa6a8af25d2..81fbadfe29fce 100644 --- a/Mathlib/LinearAlgebra/Dimension/Constructions.lean +++ b/Mathlib/LinearAlgebra/Dimension/Constructions.lean @@ -57,7 +57,7 @@ theorem LinearIndependent.sum_elim_of_quotient theorem LinearIndependent.union_of_quotient {M' : Submodule R M} {s : Set M} (hs : s ⊆ M') (hs' : LinearIndependent (ι := s) R Subtype.val) {t : Set M} (ht : LinearIndependent (ι := t) R (Submodule.Quotient.mk (p := M') ∘ Subtype.val)) : - LinearIndependent (ι := (s ∪ t : _)) R Subtype.val := by + LinearIndependent (ι := (s ∪ t :)) R Subtype.val := by refine (LinearIndependent.sum_elim_of_quotient (f := Set.embeddingOfSubset s M' hs) (of_comp M'.subtype (by simpa using hs')) Subtype.val ht).to_subtype_range' ?_ simp only [embeddingOfSubset_apply_coe, Sum.elim_range, Subtype.range_val] diff --git a/Mathlib/LinearAlgebra/FiniteDimensional.lean b/Mathlib/LinearAlgebra/FiniteDimensional.lean index 3795257b53d17..1a9987e8e47be 100644 --- a/Mathlib/LinearAlgebra/FiniteDimensional.lean +++ b/Mathlib/LinearAlgebra/FiniteDimensional.lean @@ -227,10 +227,6 @@ theorem LinearIndependent.span_eq_top_of_card_eq_finrank {ι : Type*} [Nonempty have : FiniteDimensional K V := .of_finrank_pos <| card_eq ▸ Fintype.card_pos lin_ind.span_eq_top_of_card_eq_finrank' card_eq -@[deprecated (since := "2024-02-14")] -alias span_eq_top_of_linearIndependent_of_card_eq_finrank := - LinearIndependent.span_eq_top_of_card_eq_finrank - /-- A linear independent family of `finrank K V` vectors forms a basis. -/ @[simps! repr_apply] noncomputable def basisOfLinearIndependentOfCardEqFinrank {ι : Type*} [Nonempty ι] [Fintype ι] diff --git a/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean b/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean index 16ac3a373d3e2..66b253e50a6c7 100644 --- a/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean +++ b/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean @@ -322,21 +322,12 @@ theorem FiniteDimensional.of_rank_eq_nat {n : ℕ} (h : Module.rank K V = n) : FiniteDimensional K V := Module.finite_of_rank_eq_nat h -@[deprecated (since := "2024-02-02")] -alias finiteDimensional_of_rank_eq_nat := FiniteDimensional.of_rank_eq_nat - theorem FiniteDimensional.of_rank_eq_zero (h : Module.rank K V = 0) : FiniteDimensional K V := Module.finite_of_rank_eq_zero h -@[deprecated (since := "2024-02-02")] -alias finiteDimensional_of_rank_eq_zero := FiniteDimensional.of_rank_eq_zero - theorem FiniteDimensional.of_rank_eq_one (h : Module.rank K V = 1) : FiniteDimensional K V := Module.finite_of_rank_eq_one h -@[deprecated (since := "2024-02-02")] -alias finiteDimensional_of_rank_eq_one := FiniteDimensional.of_rank_eq_one - variable (K V) instance finiteDimensional_bot : FiniteDimensional K (⊥ : Submodule K V) := diff --git a/Mathlib/LinearAlgebra/Finsupp/LinearCombination.lean b/Mathlib/LinearAlgebra/Finsupp/LinearCombination.lean index 24ebf16e2af46..c1b6ee8d30eb5 100644 --- a/Mathlib/LinearAlgebra/Finsupp/LinearCombination.lean +++ b/Mathlib/LinearAlgebra/Finsupp/LinearCombination.lean @@ -283,6 +283,11 @@ theorem linearCombinationOn_range (s : Set α) : @[deprecated (since := "2024-08-29")] alias totalOn_range := linearCombinationOn_range +theorem linearCombination_restrict (s : Set α) : + linearCombination R (s.restrict v) = Submodule.subtype _ ∘ₗ + linearCombinationOn α M R v s ∘ₗ (supportedEquivFinsupp s).symm.toLinearMap := by + classical ext; simp [linearCombinationOn] + theorem linearCombination_comp (f : α' → α) : linearCombination R (v ∘ f) = (linearCombination R v).comp (lmapDomain R R f) := by ext diff --git a/Mathlib/LinearAlgebra/Finsupp/Supported.lean b/Mathlib/LinearAlgebra/Finsupp/Supported.lean index db7332ff5630f..d779324d7594f 100644 --- a/Mathlib/LinearAlgebra/Finsupp/Supported.lean +++ b/Mathlib/LinearAlgebra/Finsupp/Supported.lean @@ -168,7 +168,7 @@ theorem disjoint_supported_supported_iff [Nontrivial M] {s t : Set α} : /-- Interpret `Finsupp.restrictSupportEquiv` as a linear equivalence between `supported M R s` and `s →₀ M`. -/ -def supportedEquivFinsupp (s : Set α) : supported M R s ≃ₗ[R] s →₀ M := by +@[simps!] def supportedEquivFinsupp (s : Set α) : supported M R s ≃ₗ[R] s →₀ M := by let F : supported M R s ≃ (s →₀ M) := restrictSupportEquiv s M refine F.toLinearEquiv ?_ have : @@ -178,6 +178,10 @@ def supportedEquivFinsupp (s : Set α) : supported M R s ≃ₗ[R] s →₀ M := rw [this] exact LinearMap.isLinear _ +@[simp] theorem supportedEquivFinsupp_symm_apply_coe (s : Set α) [DecidablePred (· ∈ s)] + (f : s →₀ M) : (supportedEquivFinsupp (R := R) s).symm f = f.extendDomain := by + convert restrictSupportEquiv_symm_apply_coe .. + section LMapDomain variable {α' : Type*} {α'' : Type*} (M R) diff --git a/Mathlib/LinearAlgebra/FreeProduct/Basic.lean b/Mathlib/LinearAlgebra/FreeProduct/Basic.lean index a6d1297916a19..74441ef678e10 100644 --- a/Mathlib/LinearAlgebra/FreeProduct/Basic.lean +++ b/Mathlib/LinearAlgebra/FreeProduct/Basic.lean @@ -64,6 +64,8 @@ end DirectSum namespace RingQuot universe uS uA uB +open scoped Function -- required for scoped `on` notation + /--If two `R`-algebras are `R`-equivalent and their quotients by a relation `rel` are defined, then their quotients are also `R`-equivalent. @@ -135,6 +137,8 @@ inductive rel : FreeTensorAlgebra R A → FreeTensorAlgebra R A → Prop (tprod R (⨁ i, A i) 2 (fun | 0 => lof R I A i a₁ | 1 => lof R I A i a₂)) (ι R <| lof R I A i (a₁ * a₂)) +open scoped Function -- required for scoped `on` notation + /--The generating equivalence relation for elements of the power algebra that are identified in the free product. -/ @[reducible, simp] def rel' := rel R A on ofDirectSum @@ -189,7 +193,7 @@ theorem mul_injections (a₁ a₂ : A i) : ι' R A (DirectSum.lof R I A i a₁) * ι' R A (DirectSum.lof R I A i a₂) = ι' R A (DirectSum.lof R I A i (a₁ * a₂)) := by convert RingQuot.mkAlgHom_rel R <| rel.prod - aesop + simp /--The `i`th canonical injection, from `A i` to the free product, as a linear map.-/ diff --git a/Mathlib/LinearAlgebra/InvariantBasisNumber.lean b/Mathlib/LinearAlgebra/InvariantBasisNumber.lean index fd7f09efc5db7..2051aaf53b365 100644 --- a/Mathlib/LinearAlgebra/InvariantBasisNumber.lean +++ b/Mathlib/LinearAlgebra/InvariantBasisNumber.lean @@ -241,8 +241,8 @@ theorem nontrivial_of_invariantBasisNumber : Nontrivial R := by exact { toFun := 0 invFun := 0 - map_add' := by aesop - map_smul' := by aesop + map_add' := by simp + map_smul' := by simp left_inv := fun _ => by simp [eq_iff_true_of_subsingleton] right_inv := fun _ => by simp [eq_iff_true_of_subsingleton] } @@ -280,8 +280,8 @@ variable {R : Type u} [CommRing R] (I : Ideal R) {ι : Type v} [Fintype ι] {ι' /-- An `R`-linear map `R^n → R^m` induces a function `R^n/I^n → R^m/I^m`. -/ private def induced_map (I : Ideal R) (e : (ι → R) →ₗ[R] ι' → R) : - (ι → R) ⧸ I.pi ι → (ι' → R) ⧸ I.pi ι' := fun x => - Quotient.liftOn' x (fun y => Ideal.Quotient.mk (I.pi ι') (e y)) + (ι → R) ⧸ Ideal.pi (fun _ ↦ I) → (ι' → R) ⧸ Ideal.pi fun _ ↦ I := fun x => + Quotient.liftOn' x (fun y => Ideal.Quotient.mk _ (e y)) (by refine fun a b hab => Ideal.Quotient.eq.2 fun h => ?_ rw [Submodule.quotientRel_def] at hab @@ -291,30 +291,14 @@ private def induced_map (I : Ideal R) (e : (ι → R) →ₗ[R] ι' → R) : /-- An isomorphism of `R`-modules `R^n ≃ R^m` induces an isomorphism of `R/I`-modules `R^n/I^n ≃ R^m/I^m`. -/ private def induced_equiv [Fintype ι'] (I : Ideal R) (e : (ι → R) ≃ₗ[R] ι' → R) : - ((ι → R) ⧸ I.pi ι) ≃ₗ[R ⧸ I] (ι' → R) ⧸ I.pi ι' where + ((ι → R) ⧸ Ideal.pi fun _ ↦ I) ≃ₗ[R ⧸ I] (ι' → R) ⧸ Ideal.pi fun _ ↦ I where -- Porting note: Lean couldn't correctly infer `(I.pi ι)` and `(I.pi ι')` on their own toFun := induced_map I e invFun := induced_map I e.symm - map_add' := by - rintro ⟨a⟩ ⟨b⟩ - convert_to Ideal.Quotient.mk (I.pi ι') _ = Ideal.Quotient.mk (I.pi ι') _ - congr - simp only [map_add] - map_smul' := by - rintro ⟨a⟩ ⟨b⟩ - convert_to Ideal.Quotient.mk (I.pi ι') _ = Ideal.Quotient.mk (I.pi ι') _ - congr - simp only [LinearEquiv.coe_coe, LinearEquiv.map_smulₛₗ, RingHom.id_apply] - left_inv := by - rintro ⟨a⟩ - convert_to Ideal.Quotient.mk (I.pi ι) _ = Ideal.Quotient.mk (I.pi ι) _ - congr - simp only [LinearEquiv.coe_coe, LinearEquiv.symm_apply_apply] - right_inv := by - rintro ⟨a⟩ - convert_to Ideal.Quotient.mk (I.pi ι') _ = Ideal.Quotient.mk (I.pi ι') _ - congr - simp only [LinearEquiv.coe_coe, LinearEquiv.apply_symm_apply] + map_add' := by rintro ⟨a⟩ ⟨b⟩; exact congr_arg _ (map_add ..) + map_smul' := by rintro ⟨a⟩ ⟨b⟩; exact congr_arg _ (map_smul ..) + left_inv := by rintro ⟨a⟩; exact congr_arg _ (e.left_inv ..) + right_inv := by rintro ⟨a⟩; exact congr_arg _ (e.right_inv ..) end @@ -332,6 +316,6 @@ instance (priority := 100) invariantBasisNumber_of_nontrivial_of_commRing {R : T ⟨fun e => let ⟨I, _hI⟩ := Ideal.exists_maximal R eq_of_fin_equiv (R ⧸ I) - ((Ideal.piQuotEquiv _ _).symm ≪≫ₗ (induced_equiv _ e ≪≫ₗ Ideal.piQuotEquiv _ _))⟩ + ((Ideal.piQuotEquiv _ _).symm ≪≫ₗ induced_equiv _ e ≪≫ₗ Ideal.piQuotEquiv _ _)⟩ end diff --git a/Mathlib/LinearAlgebra/LinearIndependent.lean b/Mathlib/LinearAlgebra/LinearIndependent.lean index 973a47eec399b..81a9e0d48fe9e 100644 --- a/Mathlib/LinearAlgebra/LinearIndependent.lean +++ b/Mathlib/LinearAlgebra/LinearIndependent.lean @@ -4,8 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro, Alexander Bentkamp, Anne Baanen -/ import Mathlib.Algebra.BigOperators.Fin +import Mathlib.Data.Fin.Tuple.Reflection import Mathlib.Data.Set.Subsingleton import Mathlib.Lean.Expr.ExtraRecognizers +import Mathlib.LinearAlgebra.Finsupp.SumProd import Mathlib.LinearAlgebra.Prod import Mathlib.Tactic.FinCases import Mathlib.Tactic.LinearCombination @@ -101,7 +103,7 @@ variable [Module R M] [Module R M'] variable (R) (v) /-- `LinearIndependent R v` states the family of vectors `v` is linearly independent over `R`. -/ def LinearIndependent : Prop := - Function.Injective (Finsupp.linearCombination R v) + Injective (Finsupp.linearCombination R v) open Lean PrettyPrinter.Delaborator SubExpr in /-- Delaborator for `LinearIndependent` that suggests pretty printing with type hints @@ -126,21 +128,27 @@ def delabLinearIndependent : Delab := variable {R v} theorem linearIndependent_iff_injective_linearCombination : - LinearIndependent R v ↔ Function.Injective (Finsupp.linearCombination R v) := Iff.rfl + LinearIndependent R v ↔ Injective (Finsupp.linearCombination R v) := Iff.rfl alias ⟨LinearIndependent.injective_linearCombination, _⟩ := linearIndependent_iff_injective_linearCombination +theorem Fintype.linearIndependent_iff_injective [Fintype ι] : + LinearIndependent R v ↔ Injective (Fintype.linearCombination R ℕ v) := by + simp [← Finsupp.linearCombination_eq_fintype_linearCombination, LinearIndependent] + +theorem LinearIndependent.injective [Nontrivial R] (hv : LinearIndependent R v) : Injective v := by + simpa [comp_def] + using Injective.comp hv (Finsupp.single_left_injective one_ne_zero) + theorem LinearIndependent.ne_zero [Nontrivial R] (i : ι) (hv : LinearIndependent R v) : v i ≠ 0 := by intro h - have := @hv (Finsupp.single i 1 : ι →₀ R) 0 ?_ - · simp at this - simpa using h - + have := @hv (Finsupp.single i 1 : ι →₀ R) 0 (by simpa using h) + simp at this theorem linearIndependent_empty_type [IsEmpty ι] : LinearIndependent R v := - Function.injective_of_subsingleton _ + injective_of_subsingleton _ variable (R M) in theorem linearIndependent_empty : LinearIndependent R (fun x => x : (∅ : Set M) → M) := @@ -150,15 +158,14 @@ theorem linearIndependent_empty : LinearIndependent R (fun x => x : (∅ : Set M linearly independent family. -/ theorem LinearIndependent.comp (h : LinearIndependent R v) (f : ι' → ι) (hf : Injective f) : LinearIndependent R (v ∘ f) := by - simpa [Function.comp_def] using Function.Injective.comp h (Finsupp.mapDomain_injective hf) + simpa [comp_def] using Injective.comp h (Finsupp.mapDomain_injective hf) /-- A set of linearly independent vectors in a module `M` over a semiring `K` is also linearly independent over a subring `R` of `K`. The implementation uses minimal assumptions about the relationship between `R`, `K` and `M`. -The version where `K` is an `R`-algebra is `LinearIndependent.restrict_scalars_algebras`. - -/ +The version where `K` is an `R`-algebra is `LinearIndependent.restrict_scalars_algebras`. -/ theorem LinearIndependent.restrict_scalars [Semiring K] [SMulWithZero R K] [Module K M] - [IsScalarTower R K M] (hinj : Function.Injective fun r : R => r • (1 : K)) + [IsScalarTower R K M] (hinj : Injective fun r : R ↦ r • (1 : K)) (li : LinearIndependent K v) : LinearIndependent R v := by intro x y hxy let f := fun r : R => r • (1 : K) @@ -167,81 +174,97 @@ theorem LinearIndependent.restrict_scalars [Semiring K] [SMulWithZero R K] [Modu exact hinj congr($this i) simpa [Finsupp.linearCombination, f, Finsupp.sum_mapRange_index] -end Semiring - -section Module - -variable {v : ι → M} -variable [Ring R] [AddCommGroup M] [AddCommGroup M'] -variable [Module R M] [Module R M'] +@[deprecated (since := "2024-08-29")] alias linearIndependent_iff_injective_total := + linearIndependent_iff_injective_linearCombination -theorem linearIndependent_iff_ker : - LinearIndependent R v ↔ LinearMap.ker (Finsupp.linearCombination R v) = ⊥ := - LinearMap.ker_eq_bot.symm +@[deprecated (since := "2024-08-29")] alias LinearIndependent.injective_total := + LinearIndependent.injective_linearCombination -theorem linearIndependent_iff : - LinearIndependent R v ↔ ∀ l, Finsupp.linearCombination R v l = 0 → l = 0 := by - simp [linearIndependent_iff_ker, LinearMap.ker_eq_bot'] +-- This version makes `l₁` and `l₂` explicit. +theorem linearIndependent_iffₛ : + LinearIndependent R v ↔ + ∀ l₁ l₂, Finsupp.linearCombination R v l₁ = Finsupp.linearCombination R v l₂ → l₁ = l₂ := + Iff.rfl -theorem linearIndependent_iff' : +open Finset in +theorem linearIndependent_iff'ₛ : LinearIndependent R v ↔ - ∀ s : Finset ι, ∀ g : ι → R, ∑ i ∈ s, g i • v i = 0 → ∀ i ∈ s, g i = 0 := - linearIndependent_iff.trans - ⟨fun hf s g hg i his => + ∀ s : Finset ι, ∀ f g : ι → R, ∑ i ∈ s, f i • v i = ∑ i ∈ s, g i • v i → ∀ i ∈ s, f i = g i := + linearIndependent_iffₛ.trans + ⟨fun hv s f g eq i his ↦ by have h := - hf (∑ i ∈ s, Finsupp.single i (g i)) <| by - simpa only [map_sum, Finsupp.linearCombination_single] using hg - calc - g i = (Finsupp.lapply i : (ι →₀ R) →ₗ[R] R) (Finsupp.single i (g i)) := by - { rw [Finsupp.lapply_apply, Finsupp.single_eq_same] } - _ = ∑ j ∈ s, (Finsupp.lapply i : (ι →₀ R) →ₗ[R] R) (Finsupp.single j (g j)) := - Eq.symm <| - Finset.sum_eq_single i - (fun j _hjs hji => by rw [Finsupp.lapply_apply, Finsupp.single_eq_of_ne hji]) - fun hnis => hnis.elim his - _ = (∑ j ∈ s, Finsupp.single j (g j)) i := (map_sum ..).symm - _ = 0 := DFunLike.ext_iff.1 h i, - fun hf _ hl => - Finsupp.ext fun _ => - _root_.by_contradiction fun hni => hni <| hf _ _ hl _ <| Finsupp.mem_support_iff.2 hni⟩ - -theorem linearIndependent_iff'' : + hv (∑ i ∈ s, Finsupp.single i (f i)) (∑ i ∈ s, Finsupp.single i (g i)) <| by + simpa only [map_sum, Finsupp.linearCombination_single] using eq + have (f : ι → R) : f i = (∑ j ∈ s, Finsupp.single j (f j)) i := + calc + f i = (Finsupp.lapply i : (ι →₀ R) →ₗ[R] R) (Finsupp.single i (f i)) := by + { rw [Finsupp.lapply_apply, Finsupp.single_eq_same] } + _ = ∑ j ∈ s, (Finsupp.lapply i : (ι →₀ R) →ₗ[R] R) (Finsupp.single j (f j)) := + Eq.symm <| + Finset.sum_eq_single i + (fun j _hjs hji => by rw [Finsupp.lapply_apply, Finsupp.single_eq_of_ne hji]) + fun hnis => hnis.elim his + _ = (∑ j ∈ s, Finsupp.single j (f j)) i := (map_sum ..).symm + rw [this f, this g, h], + fun hv f g hl ↦ + Finsupp.ext fun _ ↦ by + classical + refine _root_.by_contradiction fun hni ↦ hni <| hv (f.support ∪ g.support) f g ?_ _ ?_ + · rwa [← sum_subset subset_union_left, ← sum_subset subset_union_right] <;> + rintro i - hi <;> rw [Finsupp.not_mem_support_iff.mp hi, zero_smul] + · contrapose! hni + simp_rw [not_mem_union, Finsupp.not_mem_support_iff] at hni + rw [hni.1, hni.2]⟩ + +theorem linearIndependent_iff''ₛ : LinearIndependent R v ↔ - ∀ (s : Finset ι) (g : ι → R), (∀ i ∉ s, g i = 0) → - ∑ i ∈ s, g i • v i = 0 → ∀ i, g i = 0 := by + ∀ (s : Finset ι) (f g : ι → R), (∀ i ∉ s, f i = g i) → + ∑ i ∈ s, f i • v i = ∑ i ∈ s, g i • v i → ∀ i, f i = g i := by classical - exact linearIndependent_iff'.trans - ⟨fun H s g hg hv i => if his : i ∈ s then H s g hv i his else hg i his, fun H s g hg i hi => by + exact linearIndependent_iff'ₛ.trans + ⟨fun H s f g eq hv i ↦ if his : i ∈ s then H s f g hv i his else eq i his, + fun H s f g eq i hi ↦ by convert - H s (fun j => if j ∈ s then g j else 0) (fun j hj => if_neg hj) - (by simp_rw [ite_smul, zero_smul, Finset.sum_extend_by_zero, hg]) i + H s (fun j ↦ if j ∈ s then f j else 0) (fun j ↦ if j ∈ s then g j else 0) + (fun j hj ↦ (if_neg hj).trans (if_neg hj).symm) + (by simp_rw [ite_smul, zero_smul, Finset.sum_extend_by_zero, eq]) i <;> exact (if_pos hi).symm⟩ -theorem not_linearIndependent_iff : - ¬LinearIndependent R v ↔ - ∃ s : Finset ι, ∃ g : ι → R, ∑ i ∈ s, g i • v i = 0 ∧ ∃ i ∈ s, g i ≠ 0 := by - rw [linearIndependent_iff'] +theorem not_linearIndependent_iffₛ : + ¬LinearIndependent R v ↔ ∃ s : Finset ι, + ∃ f g : ι → R, ∑ i ∈ s, f i • v i = ∑ i ∈ s, g i • v i ∧ ∃ i ∈ s, f i ≠ g i := by + rw [linearIndependent_iff'ₛ] simp only [exists_prop, not_forall] -theorem Fintype.linearIndependent_iff [Fintype ι] : - LinearIndependent R v ↔ ∀ g : ι → R, ∑ i, g i • v i = 0 → ∀ i, g i = 0 := by - refine - ⟨fun H g => by simpa using linearIndependent_iff'.1 H Finset.univ g, fun H => - linearIndependent_iff''.2 fun s g hg hs i => H _ ?_ _⟩ - rw [← hs] - refine (Finset.sum_subset (Finset.subset_univ _) fun i _ hi => ?_).symm - rw [hg i hi, zero_smul] +theorem Fintype.linearIndependent_iffₛ [Fintype ι] : + LinearIndependent R v ↔ ∀ f g : ι → R, ∑ i, f i • v i = ∑ i, g i • v i → ∀ i, f i = g i := by + simp_rw [Fintype.linearIndependent_iff_injective, + Injective, Fintype.linearCombination_apply, funext_iff] /-- A finite family of vectors `v i` is linear independent iff the linear map that sends -`c : ι → R` to `∑ i, c i • v i` has the trivial kernel. -/ -theorem Fintype.linearIndependent_iff' [Fintype ι] [DecidableEq ι] : +`c : ι → R` to `∑ i, c i • v i` is injective. -/ +theorem Fintype.linearIndependent_iff'ₛ [Fintype ι] [DecidableEq ι] : LinearIndependent R v ↔ - LinearMap.ker (LinearMap.lsum R (fun _ ↦ R) ℕ fun i ↦ LinearMap.id.smulRight (v i)) = ⊥ := by - simp [Fintype.linearIndependent_iff, LinearMap.ker_eq_bot', funext_iff] + Injective (LinearMap.lsum R (fun _ ↦ R) ℕ fun i ↦ LinearMap.id.smulRight (v i)) := by + simp [Fintype.linearIndependent_iffₛ, Injective, funext_iff] -theorem Fintype.not_linearIndependent_iff [Fintype ι] : - ¬LinearIndependent R v ↔ ∃ g : ι → R, ∑ i, g i • v i = 0 ∧ ∃ i, g i ≠ 0 := by - simpa using not_iff_not.2 Fintype.linearIndependent_iff +theorem Fintype.not_linearIndependent_iffₛ [Fintype ι] : + ¬LinearIndependent R v ↔ ∃ f g : ι → R, ∑ i, f i • v i = ∑ i, g i • v i ∧ ∃ i, f i ≠ g i := by + simpa using not_iff_not.2 Fintype.linearIndependent_iffₛ + +lemma LinearIndependent.pair_iffₛ {x y : M} : + LinearIndependent R ![x, y] ↔ + ∀ (s t s' t' : R), s • x + t • y = s' • x + t' • y → s = s' ∧ t = t' := by + simp [Fintype.linearIndependent_iffₛ, Fin.forall_fin_two, ← FinVec.forall_iff]; rfl + +lemma LinearIndependent.eq_of_pair {x y : M} (h : LinearIndependent R ![x, y]) + {s t s' t' : R} (h' : s • x + t • y = s' • x + t' • y) : s = s' ∧ t = t' := + pair_iffₛ.mp h _ _ _ _ h' + +lemma LinearIndependent.eq_zero_of_pair' {x y : M} (h : LinearIndependent R ![x, y]) + {s t : R} (h' : s • x = t • y) : s = 0 ∧ t = 0 := by + suffices H : s = 0 ∧ 0 = t from ⟨H.1, H.2.symm⟩ + exact h.eq_of_pair (by simpa using h') lemma LinearIndependent.eq_zero_of_pair {x y : M} (h : LinearIndependent R ![x, y]) {s t : R} (h' : s • x + t • y = 0) : s = 0 ∧ t = 0 := by @@ -249,107 +272,75 @@ lemma LinearIndependent.eq_zero_of_pair {x y : M} (h : LinearIndependent R ![x, · exact ⟨by simpa using congr($h 0), by simpa using congr($h 1)⟩ simpa -/-- Also see `LinearIndependent.pair_iff'` for a simpler version over fields. -/ -lemma LinearIndependent.pair_iff {x y : M} : - LinearIndependent R ![x, y] ↔ ∀ (s t : R), s • x + t • y = 0 → s = 0 ∧ t = 0 := by - refine ⟨fun h s t hst ↦ h.eq_zero_of_pair hst, fun h ↦ ?_⟩ - apply Fintype.linearIndependent_iff.2 - intro g hg - simp only [Fin.sum_univ_two, Matrix.cons_val_zero, Matrix.cons_val_one, Matrix.head_cons] at hg - intro i - fin_cases i - exacts [(h _ _ hg).1, (h _ _ hg).2] - /-- A family is linearly independent if and only if all of its finite subfamily is linearly independent. -/ theorem linearIndependent_iff_finset_linearIndependent : LinearIndependent R v ↔ ∀ (s : Finset ι), LinearIndependent R (v ∘ (Subtype.val : s → ι)) := - ⟨fun H _ ↦ H.comp _ Subtype.val_injective, fun H ↦ linearIndependent_iff'.2 fun s g hg i hi ↦ - Fintype.linearIndependent_iff.1 (H s) (g ∘ Subtype.val) - (hg ▸ Finset.sum_attach s fun j ↦ g j • v j) ⟨i, hi⟩⟩ + ⟨fun H _ ↦ H.comp _ Subtype.val_injective, fun H ↦ linearIndependent_iff'ₛ.2 fun s f g eq i hi ↦ + Fintype.linearIndependent_iffₛ.1 (H s) (f ∘ Subtype.val) (g ∘ Subtype.val) + (by simpa only [← s.sum_coe_sort] using eq) ⟨i, hi⟩⟩ theorem LinearIndependent.coe_range (i : LinearIndependent R v) : LinearIndependent R ((↑) : range v → M) := by simpa using i.comp _ (rangeSplitting_injective v) -/-- If `v` is a linearly independent family of vectors and the kernel of a linear map `f` is -disjoint with the submodule spanned by the vectors of `v`, then `f ∘ v` is a linearly independent -family of vectors. See also `LinearIndependent.map'` for a special case assuming `ker f = ⊥`. -/ -theorem LinearIndependent.map (hv : LinearIndependent R v) {f : M →ₗ[R] M'} - (hf_inj : Disjoint (span R (range v)) (LinearMap.ker f)) : LinearIndependent R (f ∘ v) := by - rw [disjoint_iff_inf_le, ← Set.image_univ, Finsupp.span_image_eq_map_linearCombination, - map_inf_eq_map_inf_comap, map_le_iff_le_comap, comap_bot, Finsupp.supported_univ, top_inf_eq] - at hf_inj - rw [linearIndependent_iff_ker] at hv ⊢ - rw [hv, le_bot_iff] at hf_inj - rw [Finsupp.linearCombination_linear_comp, LinearMap.ker_comp, hf_inj] - /-- If `v` is an injective family of vectors such that `f ∘ v` is linearly independent, then `v` spans a submodule disjoint from the kernel of `f` -/ theorem Submodule.range_ker_disjoint {f : M →ₗ[R] M'} (hv : LinearIndependent R (f ∘ v)) : Disjoint (span R (range v)) (LinearMap.ker f) := by - rw [linearIndependent_iff_ker, Finsupp.linearCombination_linear_comp, LinearMap.ker_comp] at hv + rw [LinearIndependent, Finsupp.linearCombination_linear_comp] at hv rw [disjoint_iff_inf_le, ← Set.image_univ, Finsupp.span_image_eq_map_linearCombination, - map_inf_eq_map_inf_comap, hv, inf_bot_eq, map_bot] - -/-- An injective linear map sends linearly independent families of vectors to linearly independent -families of vectors. See also `LinearIndependent.map` for a more general statement. -/ -theorem LinearIndependent.map' (hv : LinearIndependent R v) (f : M →ₗ[R] M') - (hf_inj : LinearMap.ker f = ⊥) : LinearIndependent R (f ∘ v) := - hv.map <| by simp [hf_inj] + map_inf_eq_map_inf_comap, (LinearMap.ker_comp _ _).symm.trans + (LinearMap.ker_eq_bot_of_injective hv), inf_bot_eq, map_bot] /-- If `M / R` and `M' / R'` are modules, `i : R' → R` is a map, `j : M →+ M'` is a monoid map, -such that they send non-zero elements to non-zero elements, and compatible with the scalar +such that they are both injective, and compatible with the scalar multiplications on `M` and `M'`, then `j` sends linearly independent families of vectors to linearly independent families of vectors. As a special case, taking `R = R'` it is `LinearIndependent.map'`. -/ -theorem LinearIndependent.map_of_injective_injective {R' : Type*} {M' : Type*} - [Ring R'] [AddCommGroup M'] [Module R' M'] (hv : LinearIndependent R v) - (i : R' → R) (j : M →+ M') (hi : ∀ r, i r = 0 → r = 0) (hj : ∀ m, j m = 0 → m = 0) +theorem LinearIndependent.map_of_injective_injectiveₛ {R' M' : Type*} + [Semiring R'] [AddCommMonoid M'] [Module R' M'] (hv : LinearIndependent R v) + (i : R' → R) (j : M →+ M') (hi : Function.Injective i) (hj : Function.Injective j) (hc : ∀ (r : R') (m : M), j (i r • m) = r • j m) : LinearIndependent R' (j ∘ v) := by - rw [linearIndependent_iff'] at hv ⊢ - intro S r' H s hs + rw [linearIndependent_iff'ₛ] at hv ⊢ + intro S r₁ r₂ H s hs simp_rw [comp_apply, ← hc, ← map_sum] at H - exact hi _ <| hv _ _ (hj _ H) s hs - -/-- If `M / R` and `M' / R'` are modules, `i : R → R'` is a surjective map which maps zero to zero, -`j : M →+ M'` is a monoid map which sends non-zero elements to non-zero elements, such that the -scalar multiplications on `M` and `M'` are compatible, then `j` sends linearly independent families -of vectors to linearly independent families of vectors. As a special case, taking `R = R'` -it is `LinearIndependent.map'`. -/ -theorem LinearIndependent.map_of_surjective_injective {R' : Type*} {M' : Type*} - [Ring R'] [AddCommGroup M'] [Module R' M'] (hv : LinearIndependent R v) - (i : ZeroHom R R') (j : M →+ M') (hi : Surjective i) (hj : ∀ m, j m = 0 → m = 0) - (hc : ∀ (r : R) (m : M), j (r • m) = i r • j m) : LinearIndependent R' (j ∘ v) := by - obtain ⟨i', hi'⟩ := hi.hasRightInverse - refine hv.map_of_injective_injective i' j (fun _ h ↦ ?_) hj fun r m ↦ ?_ - · apply_fun i at h - rwa [hi', i.map_zero] at h - rw [hc (i' r) m, hi'] + exact hi <| hv _ _ _ (hj H) s hs /-- If the image of a family of vectors under a linear map is linearly independent, then so is the original family. -/ theorem LinearIndependent.of_comp (f : M →ₗ[R] M') (hfv : LinearIndependent R (f ∘ v)) : - LinearIndependent R v := - linearIndependent_iff'.2 fun s g hg i his => - have : (∑ i ∈ s, g i • f (v i)) = 0 := by - simp_rw [← map_smul, ← map_sum, hg, f.map_zero] - linearIndependent_iff'.1 hfv s g this i his - -/-- If `f` is an injective linear map, then the family `f ∘ v` is linearly independent -if and only if the family `v` is linearly independent. -/ -protected theorem LinearMap.linearIndependent_iff (f : M →ₗ[R] M') (hf_inj : LinearMap.ker f = ⊥) : - LinearIndependent R (f ∘ v) ↔ LinearIndependent R v := - ⟨fun h => h.of_comp f, fun h => h.map <| by simp only [hf_inj, disjoint_bot_right]⟩ + LinearIndependent R v := by + rw [LinearIndependent, Finsupp.linearCombination_linear_comp, LinearMap.coe_comp] at hfv + exact hfv.of_comp + +/-- If `f` is a linear map injective on the span of the range of `v`, then the family `f ∘ v` +is linearly independent if and only if the family `v` is linearly independent. +See `LinearMap.linearIndependent_iff_of_disjoint` for the version with `Set.InjOn` replaced +by `Disjoint` when working over a ring. -/ +protected theorem LinearMap.linearIndependent_iff_of_injOn (f : M →ₗ[R] M') + (hf_inj : Set.InjOn f (span R (Set.range v))) : + LinearIndependent R (f ∘ v) ↔ LinearIndependent R v := by + simp_rw [LinearIndependent, Finsupp.linearCombination_linear_comp, coe_comp] + rw [hf_inj.injective_iff] + rw [← Finsupp.range_linearCombination, LinearMap.range_coe] + +/-- If a linear map is injective on the span of a family of linearly independent vectors, then +the family stays linearly independent after composing with the linear map. +See `LinearIndependent.map` for the version with `Set.InjOn` replaced by `Disjoint` +when working over a ring. -/ +theorem LinearIndependent.map_injOn (hv : LinearIndependent R v) (f : M →ₗ[R] M') + (hf_inj : Set.InjOn f (span R (Set.range v))) : LinearIndependent R (f ∘ v) := + (f.linearIndependent_iff_of_injOn hf_inj).mpr hv @[nontriviality] theorem linearIndependent_of_subsingleton [Subsingleton R] : LinearIndependent R v := - linearIndependent_iff.2 fun _l _hl => Subsingleton.elim _ _ + linearIndependent_iffₛ.2 fun _l _l' _hl => Subsingleton.elim _ _ theorem linearIndependent_equiv (e : ι ≃ ι') {f : ι' → M} : LinearIndependent R (f ∘ e) ↔ LinearIndependent R f := - ⟨fun h => Function.comp_id f ▸ e.self_comp_symm ▸ h.comp _ e.symm.injective, fun h => - h.comp _ e.injective⟩ + ⟨fun h ↦ comp_id f ▸ e.self_comp_symm ▸ h.comp _ e.symm.injective, + fun h ↦ h.comp _ e.injective⟩ theorem linearIndependent_equiv' (e : ι ≃ ι') {f : ι' → M} {g : ι → M} (h : f ∘ e = g) : LinearIndependent R g ↔ LinearIndependent R f := @@ -361,122 +352,129 @@ theorem linearIndependent_subtype_range {ι} {f : ι → M} (hf : Injective f) : alias ⟨LinearIndependent.of_subtype_range, _⟩ := linearIndependent_subtype_range +theorem LinearIndependent.to_subtype_range {ι} {f : ι → M} (hf : LinearIndependent R f) : + LinearIndependent R ((↑) : range f → M) := by + nontriviality R + exact (linearIndependent_subtype_range hf.injective).2 hf + +theorem LinearIndependent.to_subtype_range' {ι} {f : ι → M} (hf : LinearIndependent R f) {t} + (ht : range f = t) : LinearIndependent R ((↑) : t → M) := + ht ▸ hf.to_subtype_range + +theorem LinearIndependent.image_of_comp {ι ι'} (s : Set ι) (f : ι → ι') (g : ι' → M) + (hs : LinearIndependent R fun x : s => g (f x)) : + LinearIndependent R fun x : f '' s => g x := by + nontriviality R + have : InjOn f s := injOn_iff_injective.2 hs.injective.of_comp + exact (linearIndependent_equiv' (Equiv.Set.imageOfInjOn f s this) rfl).1 hs + +theorem LinearIndependent.image {ι} {s : Set ι} {f : ι → M} + (hs : LinearIndependent R fun x : s => f x) : + LinearIndependent R fun x : f '' s => (x : M) := by + convert LinearIndependent.image_of_comp s f id hs + +theorem LinearIndependent.group_smul {G : Type*} [hG : Group G] [DistribMulAction G R] + [DistribMulAction G M] [IsScalarTower G R M] [SMulCommClass G R M] {v : ι → M} + (hv : LinearIndependent R v) (w : ι → G) : LinearIndependent R (w • v) := by + rw [linearIndependent_iff''ₛ] at hv ⊢ + intro s g₁ g₂ hgs hsum i + refine (Group.isUnit (w i)).smul_left_cancel.mp ?_ + refine hv s (fun i ↦ w i • g₁ i) (fun i ↦ w i • g₂ i) (fun i hi ↦ ?_) ?_ i + · simp_rw [hgs i hi] + · simpa only [smul_assoc, smul_comm] using hsum + +-- This lemma cannot be proved with `LinearIndependent.group_smul` since the action of +-- `Rˣ` on `R` is not commutative. +theorem LinearIndependent.units_smul {v : ι → M} (hv : LinearIndependent R v) (w : ι → Rˣ) : + LinearIndependent R (w • v) := by + rw [linearIndependent_iff''ₛ] at hv ⊢ + intro s g₁ g₂ hgs hsum i + rw [← (w i).mul_left_inj] + refine hv s (fun i ↦ g₁ i • w i) (fun i ↦ g₂ i • w i) (fun i hi ↦ ?_) ?_ i + · simp_rw [hgs i hi] + · simpa only [smul_eq_mul, mul_smul, Pi.smul_apply'] using hsum + theorem linearIndependent_image {ι} {s : Set ι} {f : ι → M} (hf : Set.InjOn f s) : - (LinearIndependent R fun x : s => f x) ↔ LinearIndependent R fun x : f '' s => (x : M) := + (LinearIndependent R fun x : s ↦ f x) ↔ LinearIndependent R fun x : f '' s => (x : M) := linearIndependent_equiv' (Equiv.Set.imageOfInjOn _ _ hf) rfl theorem linearIndependent_span (hs : LinearIndependent R v) : LinearIndependent R (M := span R (range v)) - (fun i : ι => ⟨v i, subset_span (mem_range_self i)⟩) := + (fun i : ι ↦ ⟨v i, subset_span (mem_range_self i)⟩) := LinearIndependent.of_comp (span R (range v)).subtype hs -/-- See `LinearIndependent.fin_cons` for a family of elements in a vector space. -/ -theorem LinearIndependent.fin_cons' {m : ℕ} (x : M) (v : Fin m → M) (hli : LinearIndependent R v) - (x_ortho : ∀ (c : R) (y : Submodule.span R (Set.range v)), c • x + y = (0 : M) → c = 0) : - LinearIndependent R (Fin.cons x v : Fin m.succ → M) := by - rw [Fintype.linearIndependent_iff] at hli ⊢ - rintro g total_eq j - simp_rw [Fin.sum_univ_succ, Fin.cons_zero, Fin.cons_succ] at total_eq - have : g 0 = 0 := by - refine x_ortho (g 0) ⟨∑ i : Fin m, g i.succ • v i, ?_⟩ total_eq - exact sum_mem fun i _ => smul_mem _ _ (subset_span ⟨i, rfl⟩) - rw [this, zero_smul, zero_add] at total_eq - exact Fin.cases this (hli _ total_eq) j - /-- Every finite subset of a linearly independent set is linearly independent. -/ theorem linearIndependent_finset_map_embedding_subtype (s : Set M) (li : LinearIndependent R ((↑) : s → M)) (t : Finset s) : - LinearIndependent R ((↑) : Finset.map (Embedding.subtype s) t → M) := by - let f : t.map (Embedding.subtype s) → s := fun x => - ⟨x.1, by - obtain ⟨x, h⟩ := x - rw [Finset.mem_map] at h - obtain ⟨a, _ha, rfl⟩ := h - simp only [Subtype.coe_prop, Embedding.coe_subtype]⟩ - convert LinearIndependent.comp li f ?_ - rintro ⟨x, hx⟩ ⟨y, hy⟩ - rw [Finset.mem_map] at hx hy - obtain ⟨a, _ha, rfl⟩ := hx - obtain ⟨b, _hb, rfl⟩ := hy - simp only [f, imp_self, Subtype.mk_eq_mk] - + LinearIndependent R ((↑) : Finset.map (Embedding.subtype s) t → M) := + li.comp (fun _ ↦ ⟨_, _⟩) <| by intro; aesop section Subtype /-! The following lemmas use the subtype defined by a set in `M` as the index set `ι`. -/ - -theorem linearIndependent_comp_subtype {s : Set ι} : +theorem linearIndependent_comp_subtypeₛ {s : Set ι} : LinearIndependent R (v ∘ (↑) : s → M) ↔ - ∀ l ∈ Finsupp.supported R R s, (Finsupp.linearCombination R v) l = 0 → l = 0 := by - simp only [linearIndependent_iff, (· ∘ ·), Finsupp.mem_supported, Finsupp.linearCombination_apply, - Set.subset_def, Finset.mem_coe] - constructor - · intro h l hl₁ hl₂ - exact (Finsupp.subtypeDomain_eq_zero_iff hl₁).1 <| - h (l.subtypeDomain s) ((Finsupp.sum_subtypeDomain_index hl₁).trans hl₂) - · intro h l hl - refine Finsupp.embDomain_eq_zero.1 (h (l.embDomain <| Function.Embedding.subtype s) ?_ ?_) - · suffices ∀ i hi, ¬l ⟨i, hi⟩ = 0 → i ∈ s by simpa - intros - assumption - · rwa [Finsupp.embDomain_eq_mapDomain, Finsupp.sum_mapDomain_index] - exacts [fun _ => zero_smul _ _, fun _ _ _ => add_smul _ _ _] - -theorem linearDependent_comp_subtype' {s : Set ι} : + ∀ f ∈ Finsupp.supported R R s, ∀ g ∈ Finsupp.supported R R s, + Finsupp.linearCombination R v f = Finsupp.linearCombination R v g → f = g := by + simp only [linearIndependent_iffₛ, (· ∘ ·), Finsupp.mem_supported, + Finsupp.linearCombination_apply, Set.subset_def, Finset.mem_coe] + refine ⟨fun h l₁ h₁ l₂ h₂ eq ↦ (Finsupp.subtypeDomain_eq_iff h₁ h₂).1 <| h _ _ <| + (Finsupp.sum_subtypeDomain_index h₁).trans eq ▸ (Finsupp.sum_subtypeDomain_index h₂).symm, + fun h l₁ l₂ eq ↦ ?_⟩ + refine Finsupp.embDomain_injective (Embedding.subtype s) <| h _ ?_ _ ?_ ?_ + iterate 2 simpa using fun _ h _ ↦ h + simp_rw [Finsupp.embDomain_eq_mapDomain] + rwa [Finsupp.sum_mapDomain_index, Finsupp.sum_mapDomain_index] <;> + intros <;> simp only [zero_smul, add_smul] + +theorem linearDependent_comp_subtype'ₛ {s : Set ι} : ¬LinearIndependent R (v ∘ (↑) : s → M) ↔ - ∃ f : ι →₀ R, f ∈ Finsupp.supported R R s ∧ Finsupp.linearCombination R v f = 0 ∧ f ≠ 0 := by - simp [linearIndependent_comp_subtype, and_left_comm] + ∃ f g : ι →₀ R, f ∈ Finsupp.supported R R s ∧ g ∈ Finsupp.supported R R s ∧ + Finsupp.linearCombination R v f = Finsupp.linearCombination R v g ∧ f ≠ g := by + simp [linearIndependent_comp_subtypeₛ, and_left_comm] -/-- A version of `linearDependent_comp_subtype'` with `Finsupp.linearCombination` unfolded. -/ -theorem linearDependent_comp_subtype {s : Set ι} : +/-- A version of `linearDependent_comp_subtype'ₛ` with `Finsupp.linearCombination` unfolded. -/ +theorem linearDependent_comp_subtypeₛ {s : Set ι} : ¬LinearIndependent R (v ∘ (↑) : s → M) ↔ - ∃ f : ι →₀ R, f ∈ Finsupp.supported R R s ∧ ∑ i ∈ f.support, f i • v i = 0 ∧ f ≠ 0 := - linearDependent_comp_subtype' - -theorem linearIndependent_subtype {s : Set M} : - LinearIndependent R (fun x => x : s → M) ↔ - ∀ l ∈ Finsupp.supported R R s, (Finsupp.linearCombination R id) l = 0 → l = 0 := by - apply linearIndependent_comp_subtype (v := id) - -theorem linearIndependent_comp_subtype_disjoint {s : Set ι} : - LinearIndependent R (v ∘ (↑) : s → M) ↔ - Disjoint (Finsupp.supported R R s) (LinearMap.ker <| Finsupp.linearCombination R v) := by - rw [linearIndependent_comp_subtype, LinearMap.disjoint_ker] - -theorem linearIndependent_subtype_disjoint {s : Set M} : - LinearIndependent R (fun x => x : s → M) ↔ - Disjoint (Finsupp.supported R R s) (LinearMap.ker <| Finsupp.linearCombination R id) := by - apply linearIndependent_comp_subtype_disjoint (v := id) - -theorem linearIndependent_iff_linearCombinationOn {s : Set M} : - LinearIndependent R (fun x => x : s → M) ↔ - (LinearMap.ker <| Finsupp.linearCombinationOn M M R id s) = ⊥ := by - rw [Finsupp.linearCombinationOn, LinearMap.ker, LinearMap.comap_codRestrict, Submodule.map_bot, - comap_bot, LinearMap.ker_comp, linearIndependent_subtype_disjoint, disjoint_iff_inf_le, - ← map_comap_subtype, map_le_iff_le_comap, comap_bot, ker_subtype, le_bot_iff] - -@[deprecated (since := "2024-08-29")] alias linearIndependent_iff_totalOn := - linearIndependent_iff_linearCombinationOn + ∃ f g : ι →₀ R, f ∈ Finsupp.supported R R s ∧ g ∈ Finsupp.supported R R s ∧ + ∑ i ∈ f.support, f i • v i = ∑ i ∈ g.support, g i • v i ∧ f ≠ g := + linearDependent_comp_subtype'ₛ + +theorem linearIndependent_subtypeₛ {s : Set M} : + LinearIndependent R (fun x ↦ x : s → M) ↔ + ∀ f ∈ Finsupp.supported R R s, ∀ g ∈ Finsupp.supported R R s, + Finsupp.linearCombination R id f = Finsupp.linearCombination R id g → f = g := + linearIndependent_comp_subtypeₛ (v := id) + +theorem linearIndependent_restrict_iff {s : Set ι} : + LinearIndependent R (s.restrict v) ↔ + Injective (Finsupp.linearCombinationOn ι M R v s) := by + simp [LinearIndependent, Finsupp.linearCombination_restrict] + +theorem linearIndependent_iff_linearCombinationOnₛ {s : Set M} : + LinearIndependent R (fun x ↦ x : s → M) ↔ + Injective (Finsupp.linearCombinationOn M M R id s) := + linearIndependent_restrict_iff (v := id) theorem LinearIndependent.restrict_of_comp_subtype {s : Set ι} (hs : LinearIndependent R (v ∘ (↑) : s → M)) : LinearIndependent R (s.restrict v) := hs -theorem LinearIndependent.mono {t s : Set M} (h : t ⊆ s) : - LinearIndependent R (fun x => x : s → M) → LinearIndependent R (fun x => x : t → M) := by - simp only [linearIndependent_subtype_disjoint] - exact Disjoint.mono_left (Finsupp.supported_mono h) +theorem LinearIndependent.mono {t s : Set M} (h : t ⊆ s) + (hs : LinearIndependent R (fun x ↦ x : s → M)) : LinearIndependent R (fun x ↦ x : t → M) := + hs.comp _ (Set.inclusion_injective h) theorem linearIndependent_of_finite (s : Set M) - (H : ∀ t ⊆ s, Set.Finite t → LinearIndependent R (fun x => x : t → M)) : - LinearIndependent R (fun x => x : s → M) := - linearIndependent_subtype.2 fun l hl => - linearIndependent_subtype.1 (H _ hl (Finset.finite_toSet _)) l (Subset.refl _) + (H : ∀ t ⊆ s, Set.Finite t → LinearIndependent R (fun x ↦ x : t → M)) : + LinearIndependent R (fun x ↦ x : s → M) := + linearIndependent_subtypeₛ.2 fun f hf g hg eq ↦ + linearIndependent_subtypeₛ.1 (H _ (union_subset hf hg) <| (Finset.finite_toSet _).union <| + Finset.finite_toSet _) f Set.subset_union_left g Set.subset_union_right eq theorem linearIndependent_iUnion_of_directed {η : Type*} {s : η → Set M} (hs : Directed (· ⊆ ·) s) - (h : ∀ i, LinearIndependent R (fun x => x : s i → M)) : - LinearIndependent R (fun x => x : (⋃ i, s i) → M) := by + (h : ∀ i, LinearIndependent R (fun x ↦ x : s i → M)) : + LinearIndependent R (fun x ↦ x : (⋃ i, s i) → M) := by by_cases hη : Nonempty η · refine linearIndependent_of_finite (⋃ i, s i) fun t ht ft => ?_ rcases finite_subset_iUnion ft ht with ⟨I, fi, hI⟩ @@ -493,112 +491,213 @@ theorem linearIndependent_sUnion_of_directed {s : Set (Set M)} (hs : DirectedOn exact linearIndependent_iUnion_of_directed hs.directed_val (by simpa using h) theorem linearIndependent_biUnion_of_directed {η} {s : Set η} {t : η → Set M} - (hs : DirectedOn (t ⁻¹'o (· ⊆ ·)) s) (h : ∀ a ∈ s, LinearIndependent R (fun x => x : t a → M)) : - LinearIndependent R (fun x => x : (⋃ a ∈ s, t a) → M) := by + (hs : DirectedOn (t ⁻¹'o (· ⊆ ·)) s) (h : ∀ a ∈ s, LinearIndependent R (fun x ↦ x : t a → M)) : + LinearIndependent R (fun x ↦ x : (⋃ a ∈ s, t a) → M) := by rw [biUnion_eq_iUnion] exact linearIndependent_iUnion_of_directed (directed_comp.2 <| hs.directed_val) (by simpa using h) end Subtype -end Module - -/-! ### Properties which require `Ring R` -/ +/-- Linear independent families are injective, even if you multiply either side. -/ +theorem LinearIndependent.eq_of_smul_apply_eq_smul_apply {M : Type*} [AddCommMonoid M] [Module R M] + {v : ι → M} (li : LinearIndependent R v) (c d : R) (i j : ι) (hc : c ≠ 0) + (h : c • v i = d • v j) : i = j := by + have h_single_eq : Finsupp.single i c = Finsupp.single j d := + li <| by simpa [Finsupp.linearCombination_apply] using h + rcases (Finsupp.single_eq_single_iff ..).mp h_single_eq with (⟨H, _⟩ | ⟨hc, _⟩) + · exact H + · contradiction +section Subtype -section Module +/-! The following lemmas use the subtype defined by a set in `M` as the index set `ι`. -/ -variable {v : ι → M} -variable [Ring R] [AddCommGroup M] [AddCommGroup M'] -variable [Module R M] [Module R M'] +theorem LinearIndependent.disjoint_span_image (hv : LinearIndependent R v) {s t : Set ι} + (hs : Disjoint s t) : Disjoint (Submodule.span R <| v '' s) (Submodule.span R <| v '' t) := by + simp only [disjoint_def, Finsupp.mem_span_image_iff_linearCombination] + rintro _ ⟨l₁, hl₁, rfl⟩ ⟨l₂, hl₂, H⟩ + rw [hv.injective_linearCombination.eq_iff] at H; subst l₂ + have : l₁ = 0 := Submodule.disjoint_def.mp (Finsupp.disjoint_supported_supported hs) _ hl₁ hl₂ + simp [this] -@[deprecated (since := "2024-08-29")] alias linearIndependent_iff_injective_total := - linearIndependent_iff_injective_linearCombination +theorem LinearIndependent.not_mem_span_image [Nontrivial R] (hv : LinearIndependent R v) {s : Set ι} + {x : ι} (h : x ∉ s) : v x ∉ Submodule.span R (v '' s) := by + have h' : v x ∈ Submodule.span R (v '' {x}) := by + rw [Set.image_singleton] + exact mem_span_singleton_self (v x) + intro w + apply LinearIndependent.ne_zero x hv + refine disjoint_def.1 (hv.disjoint_span_image ?_) (v x) h' w + simpa using h -@[deprecated (since := "2024-08-29")] alias LinearIndependent.injective_total := - LinearIndependent.injective_linearCombination +theorem LinearIndependent.linearCombination_ne_of_not_mem_support [Nontrivial R] + (hv : LinearIndependent R v) {x : ι} (f : ι →₀ R) (h : x ∉ f.support) : + Finsupp.linearCombination R v f ≠ v x := by + replace h : x ∉ (f.support : Set ι) := h + have p := hv.not_mem_span_image h + intro w + rw [← w] at p + rw [Finsupp.span_image_eq_map_linearCombination] at p + simp only [not_exists, not_and, mem_map] at p -- Porting note: `mem_map` isn't currently triggered + exact p f (f.mem_supported_support R) rfl -theorem LinearIndependent.injective [Nontrivial R] (hv : LinearIndependent R v) : Injective v := by - simpa [Function.comp_def] - using Function.Injective.comp hv (Finsupp.single_left_injective one_ne_zero) +@[deprecated (since := "2024-08-29")] alias LinearIndependent.total_ne_of_not_mem_support := + LinearIndependent.linearCombination_ne_of_not_mem_support -theorem LinearIndependent.to_subtype_range {ι} {f : ι → M} (hf : LinearIndependent R f) : - LinearIndependent R ((↑) : range f → M) := by - nontriviality R - exact (linearIndependent_subtype_range hf.injective).2 hf +end Subtype -theorem LinearIndependent.to_subtype_range' {ι} {f : ι → M} (hf : LinearIndependent R f) {t} - (ht : range f = t) : LinearIndependent R ((↑) : t → M) := - ht ▸ hf.to_subtype_range +section repr -theorem LinearIndependent.image_of_comp {ι ι'} (s : Set ι) (f : ι → ι') (g : ι' → M) - (hs : LinearIndependent R fun x : s => g (f x)) : - LinearIndependent R fun x : f '' s => g x := by - nontriviality R - have : InjOn f s := injOn_iff_injective.2 hs.injective.of_comp - exact (linearIndependent_equiv' (Equiv.Set.imageOfInjOn f s this) rfl).1 hs +/-- Canonical isomorphism between linear combinations and the span of linearly independent vectors. +-/ +@[simps (config := { rhsMd := default }) symm_apply] +def LinearIndependent.linearCombinationEquiv (hv : LinearIndependent R v) : + (ι →₀ R) ≃ₗ[R] span R (range v) := by + refine LinearEquiv.ofBijective (LinearMap.codRestrict (span R (range v)) + (Finsupp.linearCombination R v) ?_) ⟨hv.codRestrict _, ?_⟩ + · simp_rw [← Finsupp.range_linearCombination]; exact fun c ↦ ⟨c, rfl⟩ + rw [← LinearMap.range_eq_top, LinearMap.range_eq_map, LinearMap.map_codRestrict, + ← LinearMap.range_le_iff_comap, range_subtype, Submodule.map_top, + Finsupp.range_linearCombination] -theorem LinearIndependent.image {ι} {s : Set ι} {f : ι → M} - (hs : LinearIndependent R fun x : s => f x) : - LinearIndependent R fun x : f '' s => (x : M) := by - convert LinearIndependent.image_of_comp s f id hs +@[deprecated (since := "2024-08-29")] noncomputable alias LinearIndependent.totalEquiv := + LinearIndependent.linearCombinationEquiv -theorem LinearIndependent.group_smul {G : Type*} [hG : Group G] [DistribMulAction G R] - [DistribMulAction G M] [IsScalarTower G R M] [SMulCommClass G R M] {v : ι → M} - (hv : LinearIndependent R v) (w : ι → G) : LinearIndependent R (w • v) := by - rw [linearIndependent_iff''] at hv ⊢ - intro s g hgs hsum i - refine (smul_eq_zero_iff_eq (w i)).1 ?_ - refine hv s (fun i => w i • g i) (fun i hi => ?_) ?_ i - · dsimp only - exact (hgs i hi).symm ▸ smul_zero _ - · rw [← hsum, Finset.sum_congr rfl _] - intros - dsimp - rw [smul_assoc, smul_comm] +-- Porting note: The original theorem generated by `simps` was +-- different from the theorem on Lean 3, and not simp-normal form. +@[simp] +theorem LinearIndependent.linearCombinationEquiv_apply_coe (hv : LinearIndependent R v) + (l : ι →₀ R) : hv.linearCombinationEquiv l = Finsupp.linearCombination R v l := rfl --- This lemma cannot be proved with `LinearIndependent.group_smul` since the action of --- `Rˣ` on `R` is not commutative. -theorem LinearIndependent.units_smul {v : ι → M} (hv : LinearIndependent R v) (w : ι → Rˣ) : - LinearIndependent R (w • v) := by - rw [linearIndependent_iff''] at hv ⊢ - intro s g hgs hsum i - rw [← (w i).mul_left_eq_zero] - refine hv s (fun i => g i • (w i : R)) (fun i hi => ?_) ?_ i - · dsimp only - exact (hgs i hi).symm ▸ zero_smul _ _ - · rw [← hsum, Finset.sum_congr rfl _] - intros - rw [Pi.smul_apply', smul_assoc, Units.smul_def] +@[deprecated (since := "2024-08-29")] alias LinearIndependent.totalEquiv_apply_coe := + LinearIndependent.linearCombinationEquiv_apply_coe +/-- Linear combination representing a vector in the span of linearly independent vectors. -lemma LinearIndependent.eq_of_pair {x y : M} (h : LinearIndependent R ![x, y]) - {s t s' t' : R} (h' : s • x + t • y = s' • x + t' • y) : s = s' ∧ t = t' := by - have : (s - s') • x + (t - t') • y = 0 := by - rw [← sub_eq_zero_of_eq h'] - match_scalars <;> noncomm_ring - simpa [sub_eq_zero] using h.eq_zero_of_pair this +Given a family of linearly independent vectors, we can represent any vector in their span as +a linear combination of these vectors. These are provided by this linear map. +It is simply one direction of `LinearIndependent.linearCombinationEquiv`. -/ +def LinearIndependent.repr (hv : LinearIndependent R v) : span R (range v) →ₗ[R] ι →₀ R := + hv.linearCombinationEquiv.symm -lemma LinearIndependent.eq_zero_of_pair' {x y : M} (h : LinearIndependent R ![x, y]) - {s t : R} (h' : s • x = t • y) : s = 0 ∧ t = 0 := by - suffices H : s = 0 ∧ 0 = t from ⟨H.1, H.2.symm⟩ - exact h.eq_of_pair (by simpa using h') +variable (hv : LinearIndependent R v) -/-- If two vectors `x` and `y` are linearly independent, so are their linear combinations -`a x + b y` and `c x + d y` provided the determinant `a * d - b * c` is nonzero. -/ -lemma LinearIndependent.linear_combination_pair_of_det_ne_zero {R M : Type*} [CommRing R] - [NoZeroDivisors R] [AddCommGroup M] [Module R M] - {x y : M} (h : LinearIndependent R ![x, y]) - {a b c d : R} (h' : a * d - b * c ≠ 0) : - LinearIndependent R ![a • x + b • y, c • x + d • y] := by - apply LinearIndependent.pair_iff.2 (fun s t hst ↦ ?_) - have H : (s * a + t * c) • x + (s * b + t * d) • y = 0 := by - convert hst using 1 - module - have I1 : s * a + t * c = 0 := (h.eq_zero_of_pair H).1 - have I2 : s * b + t * d = 0 := (h.eq_zero_of_pair H).2 - have J1 : (a * d - b * c) * s = 0 := by linear_combination d * I1 - c * I2 - have J2 : (a * d - b * c) * t = 0 := by linear_combination -b * I1 + a * I2 - exact ⟨by simpa [h'] using mul_eq_zero.1 J1, by simpa [h'] using mul_eq_zero.1 J2⟩ +@[simp] +theorem LinearIndependent.linearCombination_repr (x) : + Finsupp.linearCombination R v (hv.repr x) = x := + Subtype.ext_iff.1 (LinearEquiv.apply_symm_apply hv.linearCombinationEquiv x) + +@[deprecated (since := "2024-08-29")] alias LinearIndependent.total_repr := + LinearIndependent.linearCombination_repr + +theorem LinearIndependent.linearCombination_comp_repr : + (Finsupp.linearCombination R v).comp hv.repr = Submodule.subtype _ := + LinearMap.ext <| hv.linearCombination_repr + +@[deprecated (since := "2024-08-29")] alias LinearIndependent.total_comp_repr := + LinearIndependent.linearCombination_comp_repr + +theorem LinearIndependent.repr_ker : LinearMap.ker hv.repr = ⊥ := by + rw [LinearIndependent.repr, LinearEquiv.ker] + +theorem LinearIndependent.repr_range : LinearMap.range hv.repr = ⊤ := by + rw [LinearIndependent.repr, LinearEquiv.range] + +theorem LinearIndependent.repr_eq {l : ι →₀ R} {x : span R (range v)} + (eq : Finsupp.linearCombination R v l = ↑x) : hv.repr x = l := by + have : + ↑((LinearIndependent.linearCombinationEquiv hv : (ι →₀ R) →ₗ[R] span R (range v)) l) = + Finsupp.linearCombination R v l := + rfl + have : (LinearIndependent.linearCombinationEquiv hv : (ι →₀ R) →ₗ[R] span R (range v)) l = x := by + rw [eq] at this + exact Subtype.ext_iff.2 this + rw [← LinearEquiv.symm_apply_apply hv.linearCombinationEquiv l] + rw [← this] + rfl + +theorem LinearIndependent.repr_eq_single (i) (x : span R (range v)) (hx : ↑x = v i) : + hv.repr x = Finsupp.single i 1 := by + apply hv.repr_eq + simp [Finsupp.linearCombination_single, hx] + +theorem LinearIndependent.span_repr_eq [Nontrivial R] (x) : + Span.repr R (Set.range v) x = + (hv.repr x).equivMapDomain (Equiv.ofInjective _ hv.injective) := by + have p : + (Span.repr R (Set.range v) x).equivMapDomain (Equiv.ofInjective _ hv.injective).symm = + hv.repr x := by + apply (LinearIndependent.linearCombinationEquiv hv).injective + ext + simp only [LinearIndependent.linearCombinationEquiv_apply_coe, Equiv.self_comp_ofInjective_symm, + LinearIndependent.linearCombination_repr, Finsupp.linearCombination_equivMapDomain, + Span.finsupp_linearCombination_repr] + ext ⟨_, ⟨i, rfl⟩⟩ + simp [← p] + +theorem LinearIndependent.not_smul_mem_span (hv : LinearIndependent R v) (i : ι) (a : R) + (ha : a • v i ∈ span R (v '' (univ \ {i}))) : a = 0 := by + rw [Finsupp.span_image_eq_map_linearCombination, mem_map] at ha + rcases ha with ⟨l, hl, e⟩ + rw [linearIndependent_iffₛ.1 hv l (Finsupp.single i a) (by simp [e])] at hl + by_contra hn + exact (not_mem_of_mem_diff (hl <| by simp [hn])) (mem_singleton _) + +/-- See also `iSupIndep_iff_linearIndependent_of_ne_zero`. -/ +theorem LinearIndependent.iSupIndep_span_singleton (hv : LinearIndependent R v) : + iSupIndep fun i => R ∙ v i := by + refine iSupIndep_def.mp fun i => ?_ + rw [disjoint_iff_inf_le] + intro m hm + simp only [mem_inf, mem_span_singleton, iSup_subtype'] at hm + rw [← span_range_eq_iSup] at hm + obtain ⟨⟨r, rfl⟩, hm⟩ := hm + suffices r = 0 by simp [this] + apply hv.not_smul_mem_span i + convert hm + ext + simp + +@[deprecated (since := "2024-11-24")] +alias LinearIndependent.independent_span_singleton := LinearIndependent.iSupIndep_span_singleton + +end repr + +section union + +open LinearMap Finsupp + +theorem LinearIndependent.image_subtypeₛ {s : Set M} {f : M →ₗ[R] M'} + (hs : LinearIndependent R (fun x ↦ x : s → M)) + (hf_inj : Set.InjOn f (span R s)) : + LinearIndependent R (fun x ↦ x : f '' s → M') := by + rw [← Subtype.range_coe (s := s)] at hf_inj + refine (hs.map_injOn f hf_inj).to_subtype_range' ?_ + simp [Set.range_comp f] + +theorem linearIndependent_inl_union_inr' {v : ι → M} {v' : ι' → M'} + (hv : LinearIndependent R v) (hv' : LinearIndependent R v') : + LinearIndependent R (Sum.elim (inl R M M' ∘ v) (inr R M M' ∘ v')) := by + have : linearCombination R (Sum.elim (inl R M M' ∘ v) (inr R M M' ∘ v')) = + .prodMap (linearCombination R v) (linearCombination R v') ∘ₗ + (sumFinsuppLEquivProdFinsupp R).toLinearMap := by + ext (_ | _) <;> simp [linearCombination_comapDomain] + rw [LinearIndependent, this] + simpa [LinearMap.coe_prodMap] using ⟨hv, hv'⟩ + +theorem LinearIndependent.inl_union_inr {s : Set M} {t : Set M'} + (hs : LinearIndependent R (fun x => x : s → M)) + (ht : LinearIndependent R (fun x => x : t → M')) : + LinearIndependent R (fun x => x : ↥(inl R M M' '' s ∪ inr R M M' '' t) → M × M') := by + nontriviality R + let e : s ⊕ t ≃ ↥(inl R M M' '' s ∪ inr R M M' '' t) := + .ofBijective (Sum.elim (fun i ↦ ⟨_, .inl ⟨_, i.2, rfl⟩⟩) fun i ↦ ⟨_, .inr ⟨_, i.2, rfl⟩⟩) + ⟨by rintro (_|_) (_|_) eq <;> simp [hs.ne_zero, ht.ne_zero] at eq <;> aesop, + by rintro ⟨_, ⟨_, _, rfl⟩ | ⟨_, _, rfl⟩⟩ <;> aesop⟩ + refine (linearIndependent_equiv' e ?_).mp (linearIndependent_inl_union_inr' hs ht) + ext (_ | _) <;> rfl + +end union section Maximal @@ -608,15 +707,15 @@ universe v w A linearly independent family is maximal if there is no strictly larger linearly independent family. -/ @[nolint unusedArguments] -def LinearIndependent.Maximal {ι : Type w} {R : Type u} [Ring R] {M : Type v} [AddCommGroup M] +def LinearIndependent.Maximal {ι : Type w} {R : Type u} [Semiring R] {M : Type v} [AddCommMonoid M] [Module R M] {v : ι → M} (_i : LinearIndependent R v) : Prop := ∀ (s : Set M) (_i' : LinearIndependent R ((↑) : s → M)) (_h : range v ≤ s), range v = s /-- An alternative characterization of a maximal linearly independent family, quantifying over types (in the same universe as `M`) into which the indexing family injects. -/ -theorem LinearIndependent.maximal_iff {ι : Type w} {R : Type u} [Ring R] [Nontrivial R] {M : Type v} - [AddCommGroup M] [Module R M] {v : ι → M} (i : LinearIndependent R v) : +theorem LinearIndependent.maximal_iff {ι : Type w} {R : Type u} [Semiring R] [Nontrivial R] + {M : Type v} [AddCommMonoid M] [Module R M] {v : ι → M} (i : LinearIndependent R v) : i.Maximal ↔ ∀ (κ : Type v) (w : κ → M) (_i' : LinearIndependent R w) (j : ι → κ) (_h : w ∘ j = v), Surjective j := by @@ -636,58 +735,299 @@ theorem LinearIndependent.maximal_iff {ι : Type w} {R : Type u} [Ring R] [Nontr rw [← image_univ, image_image] at q simpa using q -end Maximal +variable (R) + +theorem exists_maximal_independent' (s : ι → M) : + ∃ I : Set ι, + (LinearIndependent R fun x : I => s x) ∧ + ∀ J : Set ι, I ⊆ J → (LinearIndependent R fun x : J => s x) → I = J := by + let indep : Set ι → Prop := fun I => LinearIndependent R (s ∘ (↑) : I → M) + let X := { I : Set ι // indep I } + let r : X → X → Prop := fun I J => I.1 ⊆ J.1 + have key : ∀ c : Set X, IsChain r c → indep (⋃ (I : X) (_ : I ∈ c), I) := by + intro c hc + dsimp [indep] + rw [linearIndependent_comp_subtypeₛ] + intro f hfsupp g hgsupp hsum + rcases eq_empty_or_nonempty c with (rfl | hn) + · rw [show f = 0 by simpa using hfsupp, show g = 0 by simpa using hgsupp] + haveI : IsRefl X r := ⟨fun _ => Set.Subset.refl _⟩ + classical + obtain ⟨I, _I_mem, hI⟩ : ∃ I ∈ c, (f.support ∪ g.support : Set ι) ⊆ I := + f.support.coe_union _ ▸ hc.directedOn.exists_mem_subset_of_finset_subset_biUnion hn <| by + simpa using And.intro hfsupp hgsupp + exact linearIndependent_comp_subtypeₛ.mp I.2 f (subset_union_left.trans hI) + g (subset_union_right.trans hI) hsum + have trans : Transitive r := fun I J K => Set.Subset.trans + obtain ⟨⟨I, hli : indep I⟩, hmax : ∀ a, r ⟨I, hli⟩ a → r a ⟨I, hli⟩⟩ := + exists_maximal_of_chains_bounded + (fun c hc => ⟨⟨⋃ I ∈ c, (I : Set ι), key c hc⟩, fun I => Set.subset_biUnion_of_mem⟩) @trans + exact ⟨I, hli, fun J hsub hli => Set.Subset.antisymm hsub (hmax ⟨J, hli⟩ hsub)⟩ + +end Maximal + +theorem surjective_of_linearIndependent_of_span [Nontrivial R] (hv : LinearIndependent R v) + (f : ι' ↪ ι) (hss : range v ⊆ span R (range (v ∘ f))) : Surjective f := by + intro i + let repr : (span R (range (v ∘ f)) : Type _) → ι' →₀ R := (hv.comp f f.injective).repr + let l := (repr ⟨v i, hss (mem_range_self i)⟩).mapDomain f + have h_total_l : Finsupp.linearCombination R v l = v i := by + dsimp only [l] + rw [Finsupp.linearCombination_mapDomain] + rw [(hv.comp f f.injective).linearCombination_repr] + -- Porting note: `rfl` isn't necessary. + have h_total_eq : Finsupp.linearCombination R v l = Finsupp.linearCombination R v + (Finsupp.single i 1) := by + rw [h_total_l, Finsupp.linearCombination_single, one_smul] + have l_eq : l = _ := hv h_total_eq + dsimp only [l] at l_eq + rw [← Finsupp.embDomain_eq_mapDomain] at l_eq + rcases Finsupp.single_of_embDomain_single (repr ⟨v i, _⟩) f i (1 : R) zero_ne_one.symm l_eq with + ⟨i', hi'⟩ + use i' + exact hi'.2 + +theorem eq_of_linearIndependent_of_span_subtype [Nontrivial R] {s t : Set M} + (hs : LinearIndependent R (fun x => x : s → M)) (h : t ⊆ s) (hst : s ⊆ span R t) : s = t := by + let f : t ↪ s := + ⟨fun x => ⟨x.1, h x.2⟩, fun a b hab => Subtype.coe_injective (Subtype.mk.inj hab)⟩ + have h_surj : Surjective f := by + apply surjective_of_linearIndependent_of_span hs f _ + convert hst <;> simp [f, comp_def] + show s = t + apply Subset.antisymm _ h + intro x hx + rcases h_surj ⟨x, hx⟩ with ⟨y, hy⟩ + convert y.mem + rw [← Subtype.mk.inj hy] + +theorem le_of_span_le_span [Nontrivial R] {s t u : Set M} (hl : LinearIndependent R ((↑) : u → M)) + (hsu : s ⊆ u) (htu : t ⊆ u) (hst : span R s ≤ span R t) : s ⊆ t := by + have := + eq_of_linearIndependent_of_span_subtype (hl.mono (Set.union_subset hsu htu)) + Set.subset_union_right (Set.union_subset (Set.Subset.trans subset_span hst) subset_span) + rw [← this]; apply Set.subset_union_left + +theorem span_le_span_iff [Nontrivial R] {s t u : Set M} (hl : LinearIndependent R ((↑) : u → M)) + (hsu : s ⊆ u) (htu : t ⊆ u) : span R s ≤ span R t ↔ s ⊆ t := + ⟨le_of_span_le_span hl hsu htu, span_mono⟩ + +end Semiring + +section Module + +variable {v : ι → M} +variable [Ring R] [AddCommGroup M] [AddCommGroup M'] +variable [Module R M] [Module R M'] + +theorem linearIndependent_iff_ker : + LinearIndependent R v ↔ LinearMap.ker (Finsupp.linearCombination R v) = ⊥ := + LinearMap.ker_eq_bot.symm + +theorem linearIndependent_iff : + LinearIndependent R v ↔ ∀ l, Finsupp.linearCombination R v l = 0 → l = 0 := by + simp [linearIndependent_iff_ker, LinearMap.ker_eq_bot'] + +theorem linearIndependent_iff' : + LinearIndependent R v ↔ + ∀ s : Finset ι, ∀ g : ι → R, ∑ i ∈ s, g i • v i = 0 → ∀ i ∈ s, g i = 0 := by + rw [linearIndependent_iff'ₛ] + refine ⟨fun h s f ↦ ?_, fun h s f g ↦ ?_⟩ + · convert h s f 0; simp_rw [Pi.zero_apply, zero_smul, Finset.sum_const_zero] + · rw [← sub_eq_zero, ← Finset.sum_sub_distrib] + convert h s (f - g) using 3 <;> simp only [Pi.sub_apply, sub_smul, sub_eq_zero] + +theorem linearIndependent_iff'' : + LinearIndependent R v ↔ + ∀ (s : Finset ι) (g : ι → R), (∀ i ∉ s, g i = 0) → ∑ i ∈ s, g i • v i = 0 → ∀ i, g i = 0 := by + classical + exact linearIndependent_iff'.trans + ⟨fun H s g hg hv i => if his : i ∈ s then H s g hv i his else hg i his, fun H s g hg i hi => by + convert + H s (fun j => if j ∈ s then g j else 0) (fun j hj => if_neg hj) + (by simp_rw [ite_smul, zero_smul, Finset.sum_extend_by_zero, hg]) i + exact (if_pos hi).symm⟩ + +theorem not_linearIndependent_iff : + ¬LinearIndependent R v ↔ + ∃ s : Finset ι, ∃ g : ι → R, ∑ i ∈ s, g i • v i = 0 ∧ ∃ i ∈ s, g i ≠ 0 := by + rw [linearIndependent_iff'] + simp only [exists_prop, not_forall] + +theorem Fintype.linearIndependent_iff [Fintype ι] : + LinearIndependent R v ↔ ∀ g : ι → R, ∑ i, g i • v i = 0 → ∀ i, g i = 0 := by + refine + ⟨fun H g => by simpa using linearIndependent_iff'.1 H Finset.univ g, fun H => + linearIndependent_iff''.2 fun s g hg hs i => H _ ?_ _⟩ + rw [← hs] + refine (Finset.sum_subset (Finset.subset_univ _) fun i _ hi => ?_).symm + rw [hg i hi, zero_smul] + +/-- A finite family of vectors `v i` is linear independent iff the linear map that sends +`c : ι → R` to `∑ i, c i • v i` has the trivial kernel. -/ +theorem Fintype.linearIndependent_iff' [Fintype ι] [DecidableEq ι] : + LinearIndependent R v ↔ + LinearMap.ker (LinearMap.lsum R (fun _ ↦ R) ℕ fun i ↦ LinearMap.id.smulRight (v i)) = ⊥ := by + simp [Fintype.linearIndependent_iff, LinearMap.ker_eq_bot', funext_iff] + +theorem Fintype.not_linearIndependent_iff [Fintype ι] : + ¬LinearIndependent R v ↔ ∃ g : ι → R, ∑ i, g i • v i = 0 ∧ ∃ i, g i ≠ 0 := by + simpa using not_iff_not.2 Fintype.linearIndependent_iff + +/-- Also see `LinearIndependent.pair_iff'` for a simpler version over fields. -/ +lemma LinearIndependent.pair_iff {x y : M} : + LinearIndependent R ![x, y] ↔ ∀ (s t : R), s • x + t • y = 0 → s = 0 ∧ t = 0 := by + refine ⟨fun h s t hst ↦ h.eq_zero_of_pair hst, fun h ↦ ?_⟩ + apply Fintype.linearIndependent_iff.2 + intro g hg + simp only [Fin.sum_univ_two, Matrix.cons_val_zero, Matrix.cons_val_one, Matrix.head_cons] at hg + intro i + fin_cases i + exacts [(h _ _ hg).1, (h _ _ hg).2] + +/-- If the kernel of a linear map is disjoint from the span of a family of vectors, +then the family is linearly independent iff it is linearly independent after composing with +the linear map. -/ +protected theorem LinearMap.linearIndependent_iff_of_disjoint (f : M →ₗ[R] M') + (hf_inj : Disjoint (span R (Set.range v)) (LinearMap.ker f)) : + LinearIndependent R (f ∘ v) ↔ LinearIndependent R v := + f.linearIndependent_iff_of_injOn <| LinearMap.injOn_of_disjoint_ker le_rfl hf_inj + +/-- If `v` is a linearly independent family of vectors and the kernel of a linear map `f` is +disjoint with the submodule spanned by the vectors of `v`, then `f ∘ v` is a linearly independent +family of vectors. See also `LinearIndependent.map'` for a special case assuming `ker f = ⊥`. -/ +theorem LinearIndependent.map (hv : LinearIndependent R v) {f : M →ₗ[R] M'} + (hf_inj : Disjoint (span R (range v)) (LinearMap.ker f)) : LinearIndependent R (f ∘ v) := + (f.linearIndependent_iff_of_disjoint hf_inj).mpr hv + +/-- An injective linear map sends linearly independent families of vectors to linearly independent +families of vectors. See also `LinearIndependent.map` for a more general statement. -/ +theorem LinearIndependent.map' (hv : LinearIndependent R v) (f : M →ₗ[R] M') + (hf_inj : LinearMap.ker f = ⊥) : LinearIndependent R (f ∘ v) := + hv.map <| by simp_rw [hf_inj, disjoint_bot_right] + +/-- If `M / R` and `M' / R'` are modules, `i : R' → R` is a map, `j : M →+ M'` is a monoid map, +such that they send non-zero elements to non-zero elements, and compatible with the scalar +multiplications on `M` and `M'`, then `j` sends linearly independent families of vectors to +linearly independent families of vectors. As a special case, taking `R = R'` +it is `LinearIndependent.map'`. -/ +theorem LinearIndependent.map_of_injective_injective {R' M' : Type*} + [Ring R'] [AddCommGroup M'] [Module R' M'] (hv : LinearIndependent R v) + (i : R' → R) (j : M →+ M') (hi : ∀ r, i r = 0 → r = 0) (hj : ∀ m, j m = 0 → m = 0) + (hc : ∀ (r : R') (m : M), j (i r • m) = r • j m) : LinearIndependent R' (j ∘ v) := by + rw [linearIndependent_iff'] at hv ⊢ + intro S r' H s hs + simp_rw [comp_apply, ← hc, ← map_sum] at H + exact hi _ <| hv _ _ (hj _ H) s hs + +/-- If `M / R` and `M' / R'` are modules, `i : R → R'` is a surjective map which maps zero to zero, +`j : M →+ M'` is a monoid map which sends non-zero elements to non-zero elements, such that the +scalar multiplications on `M` and `M'` are compatible, then `j` sends linearly independent families +of vectors to linearly independent families of vectors. As a special case, taking `R = R'` +it is `LinearIndependent.map'`. -/ +theorem LinearIndependent.map_of_surjective_injective {R' : Type*} {M' : Type*} + [Ring R'] [AddCommGroup M'] [Module R' M'] (hv : LinearIndependent R v) + (i : ZeroHom R R') (j : M →+ M') (hi : Surjective i) (hj : ∀ m, j m = 0 → m = 0) + (hc : ∀ (r : R) (m : M), j (r • m) = i r • j m) : LinearIndependent R' (j ∘ v) := by + obtain ⟨i', hi'⟩ := hi.hasRightInverse + refine hv.map_of_injective_injective i' j (fun _ h ↦ ?_) hj fun r m ↦ ?_ + · apply_fun i at h + rwa [hi', i.map_zero] at h + rw [hc (i' r) m, hi'] + +/-- If `f` is an injective linear map, then the family `f ∘ v` is linearly independent +if and only if the family `v` is linearly independent. -/ +protected theorem LinearMap.linearIndependent_iff (f : M →ₗ[R] M') (hf_inj : LinearMap.ker f = ⊥) : + LinearIndependent R (f ∘ v) ↔ LinearIndependent R v := + f.linearIndependent_iff_of_disjoint <| by simp_rw [hf_inj, disjoint_bot_right] + +/-- See `LinearIndependent.fin_cons` for a family of elements in a vector space. -/ +theorem LinearIndependent.fin_cons' {m : ℕ} (x : M) (v : Fin m → M) (hli : LinearIndependent R v) + (x_ortho : ∀ (c : R) (y : Submodule.span R (Set.range v)), c • x + y = (0 : M) → c = 0) : + LinearIndependent R (Fin.cons x v : Fin m.succ → M) := by + rw [Fintype.linearIndependent_iff] at hli ⊢ + rintro g total_eq j + simp_rw [Fin.sum_univ_succ, Fin.cons_zero, Fin.cons_succ] at total_eq + have : g 0 = 0 := by + refine x_ortho (g 0) ⟨∑ i : Fin m, g i.succ • v i, ?_⟩ total_eq + exact sum_mem fun i _ => smul_mem _ _ (subset_span ⟨i, rfl⟩) + rw [this, zero_smul, zero_add] at total_eq + exact Fin.cases this (hli _ total_eq) j + +section Subtype + +/-! The following lemmas use the subtype defined by a set in `M` as the index set `ι`. -/ + +theorem linearIndependent_comp_subtype {s : Set ι} : + LinearIndependent R (v ∘ (↑) : s → M) ↔ + ∀ l ∈ Finsupp.supported R R s, (Finsupp.linearCombination R v) l = 0 → l = 0 := + linearIndependent_comp_subtypeₛ.trans ⟨fun h l hl ↦ h l hl 0 (zero_mem _), fun h f hf g hg eq ↦ + sub_eq_zero.mp (h (f - g) (sub_mem hf hg) <| by rw [map_sub, eq, sub_self])⟩ + +theorem linearDependent_comp_subtype' {s : Set ι} : + ¬LinearIndependent R (v ∘ (↑) : s → M) ↔ + ∃ f : ι →₀ R, f ∈ Finsupp.supported R R s ∧ Finsupp.linearCombination R v f = 0 ∧ f ≠ 0 := by + simp [linearIndependent_comp_subtype, and_left_comm] + +/-- A version of `linearDependent_comp_subtype'` with `Finsupp.linearCombination` unfolded. -/ +theorem linearDependent_comp_subtype {s : Set ι} : + ¬LinearIndependent R (v ∘ (↑) : s → M) ↔ + ∃ f : ι →₀ R, f ∈ Finsupp.supported R R s ∧ ∑ i ∈ f.support, f i • v i = 0 ∧ f ≠ 0 := + linearDependent_comp_subtype' + +theorem linearIndependent_subtype {s : Set M} : + LinearIndependent R (fun x => x : s → M) ↔ + ∀ l ∈ Finsupp.supported R R s, (Finsupp.linearCombination R id) l = 0 → l = 0 := by + apply linearIndependent_comp_subtype (v := id) + +theorem linearIndependent_comp_subtype_disjoint {s : Set ι} : + LinearIndependent R (v ∘ (↑) : s → M) ↔ + Disjoint (Finsupp.supported R R s) (LinearMap.ker <| Finsupp.linearCombination R v) := by + rw [linearIndependent_comp_subtype, LinearMap.disjoint_ker] -/-- Linear independent families are injective, even if you multiply either side. -/ -theorem LinearIndependent.eq_of_smul_apply_eq_smul_apply {M : Type*} [AddCommGroup M] [Module R M] - {v : ι → M} (li : LinearIndependent R v) (c d : R) (i j : ι) (hc : c ≠ 0) - (h : c • v i = d • v j) : i = j := by - let l : ι →₀ R := Finsupp.single i c - Finsupp.single j d - have h_total : Finsupp.linearCombination R v l = 0 := by - simp_rw [l, LinearMap.map_sub, Finsupp.linearCombination_apply] - simp [h] - have h_single_eq : Finsupp.single i c = Finsupp.single j d := by - rw [linearIndependent_iff] at li - simp [eq_add_of_sub_eq' (li l h_total)] - rcases (Finsupp.single_eq_single_iff ..).mp h_single_eq with (⟨H, _⟩ | ⟨hc, _⟩) - · exact H - · contradiction +theorem linearIndependent_subtype_disjoint {s : Set M} : + LinearIndependent R (fun x ↦ x : s → M) ↔ + Disjoint (Finsupp.supported R R s) (LinearMap.ker <| Finsupp.linearCombination R id) := by + apply linearIndependent_comp_subtype_disjoint (v := id) -section Subtype +theorem linearIndependent_iff_linearCombinationOn {s : Set M} : + LinearIndependent R (fun x ↦ x : s → M) ↔ + (LinearMap.ker <| Finsupp.linearCombinationOn M M R id s) = ⊥ := + linearIndependent_iff_linearCombinationOnₛ.trans <| + LinearMap.ker_eq_bot (M := Finsupp.supported R R s).symm -/-! The following lemmas use the subtype defined by a set in `M` as the index set `ι`. -/ +@[deprecated (since := "2024-08-29")] alias linearIndependent_iff_totalOn := + linearIndependent_iff_linearCombinationOn -theorem LinearIndependent.disjoint_span_image (hv : LinearIndependent R v) {s t : Set ι} - (hs : Disjoint s t) : Disjoint (Submodule.span R <| v '' s) (Submodule.span R <| v '' t) := by - simp only [disjoint_def, Finsupp.mem_span_image_iff_linearCombination] - rintro _ ⟨l₁, hl₁, rfl⟩ ⟨l₂, hl₂, H⟩ - rw [hv.injective_linearCombination.eq_iff] at H; subst l₂ - have : l₁ = 0 := Submodule.disjoint_def.mp (Finsupp.disjoint_supported_supported hs) _ hl₁ hl₂ - simp [this] +end Subtype -theorem LinearIndependent.not_mem_span_image [Nontrivial R] (hv : LinearIndependent R v) {s : Set ι} - {x : ι} (h : x ∉ s) : v x ∉ Submodule.span R (v '' s) := by - have h' : v x ∈ Submodule.span R (v '' {x}) := by - rw [Set.image_singleton] - exact mem_span_singleton_self (v x) - intro w - apply LinearIndependent.ne_zero x hv - refine disjoint_def.1 (hv.disjoint_span_image ?_) (v x) h' w - simpa using h +end Module -theorem LinearIndependent.linearCombination_ne_of_not_mem_support [Nontrivial R] - (hv : LinearIndependent R v) {x : ι} (f : ι →₀ R) (h : x ∉ f.support) : - Finsupp.linearCombination R v f ≠ v x := by - replace h : x ∉ (f.support : Set ι) := h - have p := hv.not_mem_span_image h - intro w - rw [← w] at p - rw [Finsupp.span_image_eq_map_linearCombination] at p - simp only [not_exists, not_and, mem_map] at p -- Porting note: `mem_map` isn't currently triggered - exact p f (f.mem_supported_support R) rfl +/-! ### Properties which require `Ring R` -/ -@[deprecated (since := "2024-08-29")] alias LinearIndependent.total_ne_of_not_mem_support := - LinearIndependent.linearCombination_ne_of_not_mem_support + +section Module + +variable {v : ι → M} +variable [Ring R] [AddCommGroup M] [AddCommGroup M'] +variable [Module R M] [Module R M'] + +/-- If two vectors `x` and `y` are linearly independent, so are their linear combinations +`a x + b y` and `c x + d y` provided the determinant `a * d - b * c` is nonzero. -/ +lemma LinearIndependent.linear_combination_pair_of_det_ne_zero {R M : Type*} [CommRing R] + [NoZeroDivisors R] [AddCommGroup M] [Module R M] + {x y : M} (h : LinearIndependent R ![x, y]) + {a b c d : R} (h' : a * d - b * c ≠ 0) : + LinearIndependent R ![a • x + b • y, c • x + d • y] := by + apply LinearIndependent.pair_iff.2 (fun s t hst ↦ ?_) + have H : (s * a + t * c) • x + (s * b + t * d) • y = 0 := by + convert hst using 1 + module + have I1 : s * a + t * c = 0 := (h.eq_zero_of_pair H).1 + have I2 : s * b + t * d = 0 := (h.eq_zero_of_pair H).2 + have J1 : (a * d - b * c) * s = 0 := by linear_combination d * I1 - c * I2 + have J2 : (a * d - b * c) * t = 0 := by linear_combination -b * I1 + a * I2 + exact ⟨by simpa [h'] using mul_eq_zero.1 J1, by simpa [h'] using mul_eq_zero.1 J2⟩ theorem linearIndependent_sum {v : ι ⊕ ι' → M} : LinearIndependent R v ↔ @@ -782,110 +1122,11 @@ theorem linearIndependent_iUnion_finite {η : Type*} {ιs : η → Type*} {f : rw [range_sigma_eq_iUnion_range] apply linearIndependent_iUnion_finite_subtype (fun j => (hindep j).to_subtype_range) hd -end Subtype - -section repr - -variable (hv : LinearIndependent R v) - -/-- Canonical isomorphism between linear combinations and the span of linearly independent vectors. --/ -@[simps (config := { rhsMd := default }) symm_apply] -def LinearIndependent.linearCombinationEquiv (hv : LinearIndependent R v) : - (ι →₀ R) ≃ₗ[R] span R (range v) := by - apply LinearEquiv.ofBijective (LinearMap.codRestrict (span R (range v)) - (Finsupp.linearCombination R v) _) - constructor - · rw [← LinearMap.ker_eq_bot, LinearMap.ker_codRestrict, ← linearIndependent_iff_ker] - · apply hv - · intro l - rw [← Finsupp.range_linearCombination] - rw [LinearMap.mem_range] - apply mem_range_self l - · rw [← LinearMap.range_eq_top, LinearMap.range_eq_map, LinearMap.map_codRestrict, - ← LinearMap.range_le_iff_comap, range_subtype, Submodule.map_top] - rw [Finsupp.range_linearCombination] - -@[deprecated (since := "2024-08-29")] noncomputable alias LinearIndependent.totalEquiv := - LinearIndependent.linearCombinationEquiv - --- Porting note: The original theorem generated by `simps` was --- different from the theorem on Lean 3, and not simp-normal form. -@[simp] -theorem LinearIndependent.linearCombinationEquiv_apply_coe (hv : LinearIndependent R v) - (l : ι →₀ R) : hv.linearCombinationEquiv l = Finsupp.linearCombination R v l := rfl - -@[deprecated (since := "2024-08-29")] alias LinearIndependent.totalEquiv_apply_coe := - LinearIndependent.linearCombinationEquiv_apply_coe -/-- Linear combination representing a vector in the span of linearly independent vectors. - -Given a family of linearly independent vectors, we can represent any vector in their span as -a linear combination of these vectors. These are provided by this linear map. -It is simply one direction of `LinearIndependent.linearCombinationEquiv`. -/ -def LinearIndependent.repr (hv : LinearIndependent R v) : span R (range v) →ₗ[R] ι →₀ R := - hv.linearCombinationEquiv.symm - -@[simp] -theorem LinearIndependent.linearCombination_repr (x) : - Finsupp.linearCombination R v (hv.repr x) = x := - Subtype.ext_iff.1 (LinearEquiv.apply_symm_apply hv.linearCombinationEquiv x) - -@[deprecated (since := "2024-08-29")] alias LinearIndependent.total_repr := - LinearIndependent.linearCombination_repr - -theorem LinearIndependent.linearCombination_comp_repr : - (Finsupp.linearCombination R v).comp hv.repr = Submodule.subtype _ := - LinearMap.ext <| hv.linearCombination_repr - -@[deprecated (since := "2024-08-29")] alias LinearIndependent.total_comp_repr := - LinearIndependent.linearCombination_comp_repr - -theorem LinearIndependent.repr_ker : LinearMap.ker hv.repr = ⊥ := by - rw [LinearIndependent.repr, LinearEquiv.ker] - -theorem LinearIndependent.repr_range : LinearMap.range hv.repr = ⊤ := by - rw [LinearIndependent.repr, LinearEquiv.range] - -theorem LinearIndependent.repr_eq {l : ι →₀ R} {x : span R (range v)} - (eq : Finsupp.linearCombination R v l = ↑x) : hv.repr x = l := by - have : - ↑((LinearIndependent.linearCombinationEquiv hv : (ι →₀ R) →ₗ[R] span R (range v)) l) = - Finsupp.linearCombination R v l := - rfl - have : (LinearIndependent.linearCombinationEquiv hv : (ι →₀ R) →ₗ[R] span R (range v)) l = x := by - rw [eq] at this - exact Subtype.ext_iff.2 this - rw [← LinearEquiv.symm_apply_apply hv.linearCombinationEquiv l] - rw [← this] - rfl - -theorem LinearIndependent.repr_eq_single (i) (x : span R (range v)) (hx : ↑x = v i) : - hv.repr x = Finsupp.single i 1 := by - apply hv.repr_eq - simp [Finsupp.linearCombination_single, hx] - -theorem LinearIndependent.span_repr_eq [Nontrivial R] (x) : - Span.repr R (Set.range v) x = - (hv.repr x).equivMapDomain (Equiv.ofInjective _ hv.injective) := by - have p : - (Span.repr R (Set.range v) x).equivMapDomain (Equiv.ofInjective _ hv.injective).symm = - hv.repr x := by - apply (LinearIndependent.linearCombinationEquiv hv).injective - ext - simp only [LinearIndependent.linearCombinationEquiv_apply_coe, Equiv.self_comp_ofInjective_symm, - LinearIndependent.linearCombination_repr, Finsupp.linearCombination_equivMapDomain, - Span.finsupp_linearCombination_repr] - ext ⟨_, ⟨i, rfl⟩⟩ - simp [← p] +open LinearMap theorem linearIndependent_iff_not_smul_mem_span : LinearIndependent R v ↔ ∀ (i : ι) (a : R), a • v i ∈ span R (v '' (univ \ {i})) → a = 0 := - ⟨fun hv i a ha => by - rw [Finsupp.span_image_eq_map_linearCombination, mem_map] at ha - rcases ha with ⟨l, hl, e⟩ - rw [sub_eq_zero.1 (linearIndependent_iff.1 hv (l - Finsupp.single i a) (by simp [e]))] at hl - by_contra hn - exact (not_mem_of_mem_diff (hl <| by simp [hn])) (mem_singleton _), fun H => + ⟨fun hv ↦ hv.not_smul_mem_span, fun H => linearIndependent_iff.2 fun l hl => by ext i; simp only [Finsupp.zero_apply] by_contra hn @@ -899,50 +1140,7 @@ theorem linearIndependent_iff_not_smul_mem_span : simp [hij] · simp [hl]⟩ -/-- See also `iSupIndep_iff_linearIndependent_of_ne_zero`. -/ -theorem LinearIndependent.iSupIndep_span_singleton (hv : LinearIndependent R v) : - iSupIndep fun i => R ∙ v i := by - refine iSupIndep_def.mp fun i => ?_ - rw [disjoint_iff_inf_le] - intro m hm - simp only [mem_inf, mem_span_singleton, iSup_subtype'] at hm - rw [← span_range_eq_iSup] at hm - obtain ⟨⟨r, rfl⟩, hm⟩ := hm - suffices r = 0 by simp [this] - apply linearIndependent_iff_not_smul_mem_span.mp hv i - convert hm - ext - simp - -@[deprecated (since := "2024-11-24")] -alias LinearIndependent.independent_span_singleton := LinearIndependent.iSupIndep_span_singleton - -variable (R) - -theorem exists_maximal_independent' (s : ι → M) : - ∃ I : Set ι, - (LinearIndependent R fun x : I => s x) ∧ - ∀ J : Set ι, I ⊆ J → (LinearIndependent R fun x : J => s x) → I = J := by - let indep : Set ι → Prop := fun I => LinearIndependent R (s ∘ (↑) : I → M) - let X := { I : Set ι // indep I } - let r : X → X → Prop := fun I J => I.1 ⊆ J.1 - have key : ∀ c : Set X, IsChain r c → indep (⋃ (I : X) (_ : I ∈ c), I) := by - intro c hc - dsimp [indep] - rw [linearIndependent_comp_subtype] - intro f hsupport hsum - rcases eq_empty_or_nonempty c with (rfl | hn) - · simpa using hsupport - haveI : IsRefl X r := ⟨fun _ => Set.Subset.refl _⟩ - obtain ⟨I, _I_mem, hI⟩ : ∃ I ∈ c, (f.support : Set ι) ⊆ I := - hc.directedOn.exists_mem_subset_of_finset_subset_biUnion hn hsupport - exact linearIndependent_comp_subtype.mp I.2 f hI hsum - have trans : Transitive r := fun I J K => Set.Subset.trans - obtain ⟨⟨I, hli : indep I⟩, hmax : ∀ a, r ⟨I, hli⟩ a → r a ⟨I, hli⟩⟩ := - exists_maximal_of_chains_bounded - (fun c hc => ⟨⟨⋃ I ∈ c, (I : Set ι), key c hc⟩, fun I => Set.subset_biUnion_of_mem⟩) @trans - exact ⟨I, hli, fun J hsub hli => Set.Subset.antisymm hsub (hmax ⟨J, hli⟩ hsub)⟩ - +variable (R) in theorem exists_maximal_independent (s : ι → M) : ∃ I : Set ι, (LinearIndependent R fun x : I => s x) ∧ @@ -977,68 +1175,11 @@ theorem exists_maximal_independent (s : ι → M) : refine neg_mem (sum_mem fun c hc => smul_mem _ _ (subset_span ⟨c, ?_, rfl⟩)) exact (memJ.mp (supp_f (Finset.erase_subset _ _ hc))).resolve_left (Finset.ne_of_mem_erase hc) -end repr - -theorem surjective_of_linearIndependent_of_span [Nontrivial R] (hv : LinearIndependent R v) - (f : ι' ↪ ι) (hss : range v ⊆ span R (range (v ∘ f))) : Surjective f := by - intro i - let repr : (span R (range (v ∘ f)) : Type _) → ι' →₀ R := (hv.comp f f.injective).repr - let l := (repr ⟨v i, hss (mem_range_self i)⟩).mapDomain f - have h_total_l : Finsupp.linearCombination R v l = v i := by - dsimp only [l] - rw [Finsupp.linearCombination_mapDomain] - rw [(hv.comp f f.injective).linearCombination_repr] - -- Porting note: `rfl` isn't necessary. - have h_total_eq : Finsupp.linearCombination R v l = Finsupp.linearCombination R v - (Finsupp.single i 1) := by - rw [h_total_l, Finsupp.linearCombination_single, one_smul] - have l_eq : l = _ := hv h_total_eq - dsimp only [l] at l_eq - rw [← Finsupp.embDomain_eq_mapDomain] at l_eq - rcases Finsupp.single_of_embDomain_single (repr ⟨v i, _⟩) f i (1 : R) zero_ne_one.symm l_eq with - ⟨i', hi'⟩ - use i' - exact hi'.2 - -theorem eq_of_linearIndependent_of_span_subtype [Nontrivial R] {s t : Set M} - (hs : LinearIndependent R (fun x => x : s → M)) (h : t ⊆ s) (hst : s ⊆ span R t) : s = t := by - let f : t ↪ s := - ⟨fun x => ⟨x.1, h x.2⟩, fun a b hab => Subtype.coe_injective (Subtype.mk.inj hab)⟩ - have h_surj : Surjective f := by - apply surjective_of_linearIndependent_of_span hs f _ - convert hst <;> simp [f, comp_def] - show s = t - apply Subset.antisymm _ h - intro x hx - rcases h_surj ⟨x, hx⟩ with ⟨y, hy⟩ - convert y.mem - rw [← Subtype.mk.inj hy] - -open LinearMap - theorem LinearIndependent.image_subtype {s : Set M} {f : M →ₗ[R] M'} - (hs : LinearIndependent R (fun x => x : s → M)) + (hs : LinearIndependent R (fun x ↦ x : s → M)) (hf_inj : Disjoint (span R s) (LinearMap.ker f)) : - LinearIndependent R (fun x => x : f '' s → M') := by - rw [← Subtype.range_coe (s := s)] at hf_inj - refine (hs.map hf_inj).to_subtype_range' ?_ - simp [Set.range_comp f] - -theorem LinearIndependent.inl_union_inr {s : Set M} {t : Set M'} - (hs : LinearIndependent R (fun x => x : s → M)) - (ht : LinearIndependent R (fun x => x : t → M')) : - LinearIndependent R (fun x => x : ↥(inl R M M' '' s ∪ inr R M M' '' t) → M × M') := by - refine (hs.image_subtype ?_).union (ht.image_subtype ?_) ?_ <;> [simp; simp; skip] - -- Note: https://github.com/leanprover-community/mathlib4/pull/8386 had to change `span_image` into `span_image _` - simp only [span_image _] - simp [disjoint_iff, prod_inf_prod] - -theorem linearIndependent_inl_union_inr' {v : ι → M} {v' : ι' → M'} (hv : LinearIndependent R v) - (hv' : LinearIndependent R v') : - LinearIndependent R (Sum.elim (inl R M M' ∘ v) (inr R M M' ∘ v')) := - (hv.map' (inl R M M') ker_inl).sum_type (hv'.map' (inr R M M') ker_inr) <| by - refine isCompl_range_inl_inr.disjoint.mono ?_ ?_ <;> - simp only [span_le, range_coe, range_comp_subset_range] + LinearIndependent R (fun x ↦ x : f '' s → M') := + hs.image_subtypeₛ <| LinearMap.injOn_of_disjoint_ker le_rfl hf_inj -- See, for example, Keith Conrad's note -- @@ -1141,7 +1282,7 @@ lemma linearIndependent_algHom_toLinearMap apply LinearIndependent.of_comp (LinearMap.ltoFun K M L) exact (linearIndependent_monoidHom M L).comp (RingHom.toMonoidHom ∘ AlgHom.toRingHom) - (fun _ _ e ↦ AlgHom.ext (DFunLike.congr_fun e : _)) + (fun _ _ e ↦ AlgHom.ext (DFunLike.congr_fun e :)) lemma linearIndependent_algHom_toLinearMap' (K M L) [CommRing K] [Semiring M] [Algebra K M] [CommRing L] [IsDomain L] [Algebra K L] [NoZeroSMulDivisors K L] : @@ -1150,17 +1291,6 @@ lemma linearIndependent_algHom_toLinearMap' (K M L) [CommRing K] simp_rw [Algebra.smul_def, mul_one] exact NoZeroSMulDivisors.algebraMap_injective K L -theorem le_of_span_le_span [Nontrivial R] {s t u : Set M} (hl : LinearIndependent R ((↑) : u → M)) - (hsu : s ⊆ u) (htu : t ⊆ u) (hst : span R s ≤ span R t) : s ⊆ t := by - have := - eq_of_linearIndependent_of_span_subtype (hl.mono (Set.union_subset hsu htu)) - Set.subset_union_right (Set.union_subset (Set.Subset.trans subset_span hst) subset_span) - rw [← this]; apply Set.subset_union_left - -theorem span_le_span_iff [Nontrivial R] {s t u : Set M} (hl : LinearIndependent R ((↑) : u → M)) - (hsu : s ⊆ u) (htu : t ⊆ u) : span R s ≤ span R t ↔ s ⊆ t := - ⟨le_of_span_le_span hl hsu htu, span_mono⟩ - end Module section Nontrivial @@ -1256,7 +1386,7 @@ theorem linearIndependent_insert' {ι} {s : Set ι} {a : ι} {f : ι → V} (has linearIndependent_option] -- Porting note: `simp [(· ∘ ·), range_comp f]` → `simp [(· ∘ ·)]; erw [range_comp f ..]; simp` -- https://github.com/leanprover-community/mathlib4/issues/5164 - simp only [Function.comp_def] + simp only [comp_def] erw [range_comp f ((↑) : s → ι)] simp @@ -1357,7 +1487,7 @@ theorem exists_linearIndependent : /-- Indexed version of `exists_linearIndependent`. -/ lemma exists_linearIndependent' (v : ι → V) : - ∃ (κ : Type u') (a : κ → ι), Function.Injective a ∧ + ∃ (κ : Type u') (a : κ → ι), Injective a ∧ Submodule.span K (Set.range (v ∘ a)) = Submodule.span K (Set.range v) ∧ LinearIndependent K (v ∘ a) := by obtain ⟨t, ht, hsp, hli⟩ := exists_linearIndependent K (Set.range v) @@ -1480,3 +1610,5 @@ theorem exists_finite_card_le_of_finite_of_linearIndependent_of_span (ht : t.Fin ⟨this, by rw [← Eq]; exact Finset.card_le_card <| Finset.coe_subset.mp <| by simp [hsu]⟩ end Module + +set_option linter.style.longFile 1700 diff --git a/Mathlib/LinearAlgebra/Matrix/Charpoly/FiniteField.lean b/Mathlib/LinearAlgebra/Matrix/Charpoly/FiniteField.lean index 3eea92af5a206..703ede3a3d235 100644 --- a/Mathlib/LinearAlgebra/Matrix/Charpoly/FiniteField.lean +++ b/Mathlib/LinearAlgebra/Matrix/Charpoly/FiniteField.lean @@ -36,7 +36,7 @@ theorem FiniteField.Matrix.charpoly_pow_card {K : Type*} [Field K] [Fintype K] ( apply congr_arg det refine matPolyEquiv.injective ?_ rw [map_pow, matPolyEquiv_charmatrix, hk, sub_pow_char_pow_of_commute, ← C_pow] - · exact (id (matPolyEquiv_eq_X_pow_sub_C (p ^ k) M) : _) + · exact (id (matPolyEquiv_eq_X_pow_sub_C (p ^ k) M) :) · exact (C M).commute_X · exact congr_arg _ (Subsingleton.elim _ _) diff --git a/Mathlib/LinearAlgebra/Matrix/Charpoly/LinearMap.lean b/Mathlib/LinearAlgebra/Matrix/Charpoly/LinearMap.lean index a6138962526ff..8befb3f75b370 100644 --- a/Mathlib/LinearAlgebra/Matrix/Charpoly/LinearMap.lean +++ b/Mathlib/LinearAlgebra/Matrix/Charpoly/LinearMap.lean @@ -64,7 +64,7 @@ theorem PiToModule.fromEnd_injective (hb : Submodule.span R (Set.range b) = ⊤) obtain ⟨m, rfl⟩ : m ∈ LinearMap.range (Fintype.linearCombination R R b) := by rw [(Fintype.range_linearCombination R b).trans hb] exact Submodule.mem_top - exact (LinearMap.congr_fun e m : _) + exact (LinearMap.congr_fun e m :) section @@ -210,7 +210,7 @@ theorem LinearMap.exists_monic_and_coeff_mem_pow_and_aeval_eq_zero_of_range_le_s cases subsingleton_or_nontrivial R · exact ⟨0, Polynomial.monic_of_subsingleton _, by simp⟩ obtain ⟨s : Finset M, hs : Submodule.span R (s : Set M) = ⊤⟩ := - Module.Finite.out (R := R) (M := M) + Module.Finite.fg_top (R := R) (M := M) -- Porting note: `H` was `rfl` obtain ⟨A, H, h⟩ := Matrix.isRepresentation.toEnd_exists_mem_ideal R ((↑) : s → M) diff --git a/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean b/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean index 57b1ced0fed82..d3f50bfe9f809 100644 --- a/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean +++ b/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean @@ -644,10 +644,7 @@ theorem det_fromBlocks_zero₂₁ (A : Matrix m m R) (B : Matrix m n R) (D : Mat · simp_rw [sum_mul_sum, ← sum_product', univ_product_univ] refine sum_nbij (fun σ ↦ σ.fst.sumCongr σ.snd) ?_ ?_ ?_ ?_ · intro σ₁₂ _ - simp only - erw [Set.mem_toFinset, MonoidHom.mem_range] - use σ₁₂ - simp only [sumCongrHom_apply] + simp · intro σ₁ _ σ₂ _ dsimp only intro h diff --git a/Mathlib/LinearAlgebra/Matrix/LDL.lean b/Mathlib/LinearAlgebra/Matrix/LDL.lean index aeee211e5e6ec..7aaa5baf5130d 100644 --- a/Mathlib/LinearAlgebra/Matrix/LDL.lean +++ b/Mathlib/LinearAlgebra/Matrix/LDL.lean @@ -47,13 +47,13 @@ variable {S : Matrix n n 𝕜} [Fintype n] (hS : S.PosDef) applying Gram-Schmidt-Orthogonalization w.r.t. the inner product induced by `Sᵀ` on the standard basis vectors `Pi.basisFun`. -/ noncomputable def LDL.lowerInv : Matrix n n 𝕜 := - @gramSchmidt 𝕜 (n → 𝕜) _ (_ : _) (InnerProductSpace.ofMatrix hS.transpose) n _ _ _ + @gramSchmidt 𝕜 (n → 𝕜) _ (_ :) (InnerProductSpace.ofMatrix hS.transpose) n _ _ _ (Pi.basisFun 𝕜 n) theorem LDL.lowerInv_eq_gramSchmidtBasis : LDL.lowerInv hS = ((Pi.basisFun 𝕜 n).toMatrix - (@gramSchmidtBasis 𝕜 (n → 𝕜) _ (_ : _) (InnerProductSpace.ofMatrix hS.transpose) n _ _ _ + (@gramSchmidtBasis 𝕜 (n → 𝕜) _ (_ :) (InnerProductSpace.ofMatrix hS.transpose) n _ _ _ (Pi.basisFun 𝕜 n)))ᵀ := by letI := NormedAddCommGroup.ofMatrix hS.transpose letI := InnerProductSpace.ofMatrix hS.transpose @@ -65,13 +65,13 @@ noncomputable instance LDL.invertibleLowerInv : Invertible (LDL.lowerInv hS) := rw [LDL.lowerInv_eq_gramSchmidtBasis] haveI := Basis.invertibleToMatrix (Pi.basisFun 𝕜 n) - (@gramSchmidtBasis 𝕜 (n → 𝕜) _ (_ : _) (InnerProductSpace.ofMatrix hS.transpose) n _ _ _ + (@gramSchmidtBasis 𝕜 (n → 𝕜) _ (_ :) (InnerProductSpace.ofMatrix hS.transpose) n _ _ _ (Pi.basisFun 𝕜 n)) infer_instance theorem LDL.lowerInv_orthogonal {i j : n} (h₀ : i ≠ j) : ⟪LDL.lowerInv hS i, Sᵀ *ᵥ LDL.lowerInv hS j⟫ₑ = 0 := - @gramSchmidt_orthogonal 𝕜 _ _ (_ : _) (InnerProductSpace.ofMatrix hS.transpose) _ _ _ _ _ _ _ h₀ + @gramSchmidt_orthogonal 𝕜 _ _ (_ :) (InnerProductSpace.ofMatrix hS.transpose) _ _ _ _ _ _ _ h₀ /-- The entries of the diagonal matrix `D` of the LDL decomposition. -/ noncomputable def LDL.diagEntries : n → 𝕜 := fun i => @@ -83,7 +83,7 @@ noncomputable def LDL.diag : Matrix n n 𝕜 := theorem LDL.lowerInv_triangular {i j : n} (hij : i < j) : LDL.lowerInv hS i j = 0 := by rw [← - @gramSchmidt_triangular 𝕜 (n → 𝕜) _ (_ : _) (InnerProductSpace.ofMatrix hS.transpose) n _ _ _ + @gramSchmidt_triangular 𝕜 (n → 𝕜) _ (_ :) (InnerProductSpace.ofMatrix hS.transpose) n _ _ _ i j hij (Pi.basisFun 𝕜 n), Pi.basisFun_repr, LDL.lowerInv] diff --git a/Mathlib/LinearAlgebra/Matrix/PosDef.lean b/Mathlib/LinearAlgebra/Matrix/PosDef.lean index fb373744d8173..5e46c1956c392 100644 --- a/Mathlib/LinearAlgebra/Matrix/PosDef.lean +++ b/Mathlib/LinearAlgebra/Matrix/PosDef.lean @@ -108,9 +108,8 @@ protected theorem natCast [StarOrderedRing R] [DecidableEq n] (d : ℕ) : rw [Nat.cast_smul_eq_nsmul] exact nsmul_nonneg (dotProduct_star_self_nonneg _) _⟩ --- See note [no_index around OfNat.ofNat] protected theorem ofNat [StarOrderedRing R] [DecidableEq n] (d : ℕ) [d.AtLeastTwo] : - PosSemidef (no_index (OfNat.ofNat d) : Matrix n n R) := + PosSemidef (ofNat(d) : Matrix n n R) := .natCast d protected theorem intCast [StarOrderedRing R] [DecidableEq n] (d : ℤ) (hd : 0 ≤ d) : @@ -386,10 +385,9 @@ theorem _root_.Matrix.posDef_natCast_iff [StarOrderedRing R] [DecidableEq n] [No PosDef (d : Matrix n n R) ↔ 0 < d := posDef_diagonal_iff.trans <| by simp --- See note [no_index around OfNat.ofNat] protected theorem ofNat [StarOrderedRing R] [DecidableEq n] [NoZeroDivisors R] (d : ℕ) [d.AtLeastTwo] : - PosDef (no_index (OfNat.ofNat d) : Matrix n n R) := + PosDef (ofNat(d) : Matrix n n R) := .natCast d (NeZero.ne _) protected theorem intCast [StarOrderedRing R] [DecidableEq n] [NoZeroDivisors R] diff --git a/Mathlib/LinearAlgebra/Matrix/Symmetric.lean b/Mathlib/LinearAlgebra/Matrix/Symmetric.lean index 498e22fb3bacc..2b2be9d38540a 100644 --- a/Mathlib/LinearAlgebra/Matrix/Symmetric.lean +++ b/Mathlib/LinearAlgebra/Matrix/Symmetric.lean @@ -126,8 +126,8 @@ theorem IsSymm.fromBlocks {A : Matrix m m α} {B : Matrix m n α} {C : Matrix n theorem isSymm_fromBlocks_iff {A : Matrix m m α} {B : Matrix m n α} {C : Matrix n m α} {D : Matrix n n α} : (A.fromBlocks B C D).IsSymm ↔ A.IsSymm ∧ Bᵀ = C ∧ Cᵀ = B ∧ D.IsSymm := ⟨fun h => - ⟨(congr_arg toBlocks₁₁ h : _), (congr_arg toBlocks₂₁ h : _), (congr_arg toBlocks₁₂ h : _), - (congr_arg toBlocks₂₂ h : _)⟩, + ⟨(congr_arg toBlocks₁₁ h :), (congr_arg toBlocks₂₁ h :), (congr_arg toBlocks₁₂ h :), + (congr_arg toBlocks₂₂ h :)⟩, fun ⟨hA, hBC, _, hD⟩ => IsSymm.fromBlocks hA hBC hD⟩ end Matrix diff --git a/Mathlib/LinearAlgebra/Multilinear/Basic.lean b/Mathlib/LinearAlgebra/Multilinear/Basic.lean index 5de11b6d8fc35..227f360fbe06f 100644 --- a/Mathlib/LinearAlgebra/Multilinear/Basic.lean +++ b/Mathlib/LinearAlgebra/Multilinear/Basic.lean @@ -195,20 +195,18 @@ theorem zero_apply (m : ∀ i, M₁ i) : (0 : MultilinearMap R M₁ M₂) m = 0 section SMul -variable {R' A : Type*} [Monoid R'] [Semiring A] [∀ i, Module A (M₁ i)] [DistribMulAction R' M₂] - [Module A M₂] [SMulCommClass A R' M₂] +variable [DistribSMul S M₂] [SMulCommClass R S M₂] -instance : SMul R' (MultilinearMap A M₁ M₂) := +instance : SMul S (MultilinearMap R M₁ M₂) := ⟨fun c f => ⟨fun m => c • f m, fun m i x y => by simp [smul_add], fun l i x d => by simp [← smul_comm x c (_ : M₂)]⟩⟩ @[simp] -theorem smul_apply (f : MultilinearMap A M₁ M₂) (c : R') (m : ∀ i, M₁ i) : (c • f) m = c • f m := +theorem smul_apply (f : MultilinearMap R M₁ M₂) (c : S) (m : ∀ i, M₁ i) : (c • f) m = c • f m := rfl -theorem coe_smul (c : R') (f : MultilinearMap A M₁ M₂) : ⇑(c • f) = c • (⇑ f) := - rfl +theorem coe_smul (c : S) (f : MultilinearMap R M₁ M₂) : ⇑(c • f) = c • (⇑ f) := rfl end SMul @@ -1081,8 +1079,8 @@ map from `Π i, M₁ᵢ ⟶ M₁ᵢ'` to `M ⟶ M₂` via `(fᵢ) ↦ v ↦ g(f MultilinearMap R M₁' M₂ →ₗ[R] MultilinearMap R (fun i ↦ M₁ i →ₗ[R] M₁' i) (MultilinearMap R M₁ M₂) where toFun g := (LinearMap.applyₗ g).compMultilinearMap compLinearMapMultilinear - map_add' := by aesop - map_smul' := by aesop + map_add' := by simp + map_smul' := by simp end @@ -1555,13 +1553,11 @@ def MultilinearMap.curryRight (f : MultilinearMap R M M₂) : map_update_add' := @fun dec m i x y => by rw [Subsingleton.elim dec (by clear dec; infer_instance)]; clear dec ext z - change f (snoc (update m i (x + y)) z) = f (snoc (update m i x) z) + f (snoc (update m i y) z) - rw [snoc_update, snoc_update, snoc_update, f.map_update_add] + simp map_update_smul' := @fun dec m i c x => by rw [Subsingleton.elim dec (by clear dec; infer_instance)]; clear dec ext z - change f (snoc (update m i (c • x)) z) = c • f (snoc (update m i x) z) - rw [snoc_update, snoc_update, f.map_update_smul] + simp @[simp] theorem MultilinearMap.curryRight_apply (f : MultilinearMap R M M₂) diff --git a/Mathlib/LinearAlgebra/Pi.lean b/Mathlib/LinearAlgebra/Pi.lean index 4c1ff2962f310..62bc97d66935b 100644 --- a/Mathlib/LinearAlgebra/Pi.lean +++ b/Mathlib/LinearAlgebra/Pi.lean @@ -217,8 +217,8 @@ theorem lsum_apply (S) [AddCommMonoid M] [Module R M] [Fintype ι] [Semiring S] lsum R φ S f = ∑ i : ι, (f i).comp (proj i) := rfl @[simp high] -theorem lsum_single {ι R : Type*} [Fintype ι] [DecidableEq ι] [CommRing R] {M : ι → Type*} - [(i : ι) → AddCommGroup (M i)] [(i : ι) → Module R (M i)] : +theorem lsum_single {ι R : Type*} [Fintype ι] [DecidableEq ι] [CommSemiring R] {M : ι → Type*} + [(i : ι) → AddCommMonoid (M i)] [(i : ι) → Module R (M i)] : LinearMap.lsum R M R (LinearMap.single R M) = LinearMap.id := LinearMap.ext fun x => by simp [Finset.univ_sum_single] @@ -458,7 +458,7 @@ def piCurry {ι : Type*} {κ : ι → Type*} (α : ∀ i, κ i → Type*) rfl /-- This is `Equiv.piOptionEquivProd` as a `LinearEquiv` -/ -def piOptionEquivProd {ι : Type*} {M : Option ι → Type*} [(i : Option ι) → AddCommGroup (M i)] +def piOptionEquivProd {ι : Type*} {M : Option ι → Type*} [(i : Option ι) → AddCommMonoid (M i)] [(i : Option ι) → Module R (M i)] : ((i : Option ι) → M i) ≃ₗ[R] M none × ((i : ι) → M (some i)) := { Equiv.piOptionEquivProd with diff --git a/Mathlib/LinearAlgebra/Prod.lean b/Mathlib/LinearAlgebra/Prod.lean index 328fdca858bd5..ed5cd352af5bd 100644 --- a/Mathlib/LinearAlgebra/Prod.lean +++ b/Mathlib/LinearAlgebra/Prod.lean @@ -993,7 +993,7 @@ variable {R S G H I : Type*} [AddCommMonoid H] [Module S H] [AddCommMonoid I] [Module S I] -/-- **Vertical line test** for module homomorphisms. +/-- **Vertical line test** for linear maps. Let `f : G → H × I` be a linear (or semilinear) map to a product. Assume that `f` is surjective on the first factor and that the image of `f` intersects every "vertical line" `{(h, i) | i : I}` at @@ -1019,10 +1019,10 @@ lemma LinearMap.exists_range_eq_graph {f : G →ₛₗ[σ] H × I} (hf₁ : Surj simpa only [mem_range, Eq.comm, ZeroHom.toFun_eq_coe, AddMonoidHom.toZeroHom_coe, mem_graph_iff, coe_mk, AddHom.coe_mk, AddMonoidHom.coe_coe, Set.mem_range] using hf' x -/-- **Vertical line test** for module homomorphisms. +/-- **Vertical line test** for linear maps. Let `G ≤ H × I` be a submodule of a product of modules. Assume that `G` maps bijectively to the -first factor. Then `G` is the graph of some module homomorphism `f : H →ₗ[R] I`. -/ +first factor. Then `G` is the graph of some linear map `f : H →ₗ[R] I`. -/ lemma Submodule.exists_eq_graph {G : Submodule S (H × I)} (hf₁ : Bijective (Prod.fst ∘ G.subtype)) : ∃ f : H →ₗ[S] I, G = LinearMap.graph f := by simpa only [range_subtype] using LinearMap.exists_range_eq_graph hf₁.surjective @@ -1030,10 +1030,10 @@ lemma Submodule.exists_eq_graph {G : Submodule S (H × I)} (hf₁ : Bijective (P /-- **Line test** for module isomorphisms. -Let `f : G → H × I` be a homomorphism to a product of modules. Assume that `f` is surjective onto -both factors and that the image of `f` intersects every "vertical line" `{(h, i) | i : I}` and every -"horizontal line" `{(h, i) | h : H}` at most once. Then the image of `f` is the graph of some -module isomorphism `f' : H ≃ I`. -/ +Let `f : G → H × I` be a linear (or semilinear) map to a product of modules. Assume that `f` is +surjective onto both factors and that the image of `f` intersects every "vertical line" +`{(h, i) | i : I}` and every "horizontal line" `{(h, i) | h : H}` at most once. Then the image of +`f` is the graph of some module isomorphism `f' : H ≃ I`. -/ lemma LinearMap.exists_linearEquiv_eq_graph {f : G →ₛₗ[σ] H × I} (hf₁ : Surjective (Prod.fst ∘ f)) (hf₂ : Surjective (Prod.snd ∘ f)) (hf : ∀ g₁ g₂, (f g₁).1 = (f g₂).1 ↔ (f g₁).2 = (f g₂).2) : ∃ e : H ≃ₗ[S] I, range f = e.toLinearMap.graph := by diff --git a/Mathlib/LinearAlgebra/QuadraticForm/Isometry.lean b/Mathlib/LinearAlgebra/QuadraticForm/Isometry.lean index 4208354f24bc0..e09af538b55e1 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/Isometry.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/Isometry.lean @@ -51,7 +51,7 @@ instance instLinearMapClass : LinearMapClass (Q₁ →qᵢ Q₂) R M₁ M₂ whe theorem toLinearMap_injective : Function.Injective (Isometry.toLinearMap : (Q₁ →qᵢ Q₂) → M₁ →ₗ[R] M₂) := fun _f _g h => - DFunLike.coe_injective (congr_arg DFunLike.coe h : _) + DFunLike.coe_injective (congr_arg DFunLike.coe h :) @[ext] theorem ext ⦃f g : Q₁ →qᵢ Q₂⦄ (h : ∀ x, f x = g x) : f = g := diff --git a/Mathlib/LinearAlgebra/Quotient/Defs.lean b/Mathlib/LinearAlgebra/Quotient/Defs.lean index 46e2cbca915b8..810d76fdc6985 100644 --- a/Mathlib/LinearAlgebra/Quotient/Defs.lean +++ b/Mathlib/LinearAlgebra/Quotient/Defs.lean @@ -241,7 +241,7 @@ variable {R₂ M₂ : Type*} [Ring R₂] [AddCommGroup M₂] [Module R₂ M₂] See note [partially-applied ext lemmas]. -/ @[ext 1100] -- Porting note: increase priority so this applies before `LinearMap.ext` theorem linearMap_qext ⦃f g : M ⧸ p →ₛₗ[τ₁₂] M₂⦄ (h : f.comp p.mkQ = g.comp p.mkQ) : f = g := - LinearMap.ext fun x => Submodule.Quotient.induction_on _ x <| (LinearMap.congr_fun h : _) + LinearMap.ext fun x => Submodule.Quotient.induction_on _ x <| (LinearMap.congr_fun h :) /-- Quotienting by equal submodules gives linearly equivalent quotients. -/ def quotEquivOfEq (h : p = p') : (M ⧸ p) ≃ₗ[R] M ⧸ p' := diff --git a/Mathlib/LinearAlgebra/Ray.lean b/Mathlib/LinearAlgebra/Ray.lean index d569b5c080c98..c460672f898e7 100644 --- a/Mathlib/LinearAlgebra/Ray.lean +++ b/Mathlib/LinearAlgebra/Ray.lean @@ -327,7 +327,7 @@ theorem someVector_ne_zero (x : Module.Ray R M) : x.someVector ≠ 0 := /-- The ray of `someVector`. -/ @[simp] theorem someVector_ray (x : Module.Ray R M) : rayOfNeZero R _ x.someVector_ne_zero = x := - (congr_arg _ (Subtype.coe_eta _ _) : _).trans x.out_eq + (congr_arg _ (Subtype.coe_eta _ _) :).trans x.out_eq end Module.Ray diff --git a/Mathlib/LinearAlgebra/Semisimple.lean b/Mathlib/LinearAlgebra/Semisimple.lean index c7032628a6166..feb4bde80dc93 100644 --- a/Mathlib/LinearAlgebra/Semisimple.lean +++ b/Mathlib/LinearAlgebra/Semisimple.lean @@ -73,8 +73,7 @@ lemma isSemisimple_iff' : lemma isSemisimple_iff : f.IsSemisimple ↔ ∀ p ∈ invtSubmodule f, ∃ q ∈ invtSubmodule f, IsCompl p q := by - simp_rw [isSemisimple_iff'] - aesop + simp [isSemisimple_iff'] lemma isSemisimple_restrict_iff (p) (hp : p ∈ invtSubmodule f) : IsSemisimple (LinearMap.restrict f hp) ↔ diff --git a/Mathlib/LinearAlgebra/SesquilinearForm.lean b/Mathlib/LinearAlgebra/SesquilinearForm.lean index e4eb407fc013c..504d2acfe7e8e 100644 --- a/Mathlib/LinearAlgebra/SesquilinearForm.lean +++ b/Mathlib/LinearAlgebra/SesquilinearForm.lean @@ -65,6 +65,7 @@ theorem isOrtho_zero_right (B : M₁ →ₛₗ[I₁] M₂ →ₛₗ[I₂] M) (x) theorem isOrtho_flip {B : M₁ →ₛₗ[I₁] M₁ →ₛₗ[I₁'] M} {x y} : B.IsOrtho x y ↔ B.flip.IsOrtho y x := by simp_rw [isOrtho_def, flip_apply] +open scoped Function in -- required for scoped `on` notation /-- A set of vectors `v` is orthogonal with respect to some bilinear map `B` if and only if for all `i ≠ j`, `B (v i) (v j) = 0`. For orthogonality between two elements, use `BilinForm.isOrtho` -/ diff --git a/Mathlib/LinearAlgebra/Span/Basic.lean b/Mathlib/LinearAlgebra/Span/Basic.lean index 7ef6263250020..9da8fb60044f6 100644 --- a/Mathlib/LinearAlgebra/Span/Basic.lean +++ b/Mathlib/LinearAlgebra/Span/Basic.lean @@ -136,10 +136,8 @@ end theorem span_smul_eq_of_isUnit (s : Set M) (r : R) (hr : IsUnit r) : span R (r • s) = span R s := by apply le_antisymm · apply span_smul_le - · convert span_smul_le (r • s) ((hr.unit⁻¹ : _) : R) - rw [smul_smul] - erw [hr.unit.inv_val] - rw [one_smul] + · convert span_smul_le (r • s) ((hr.unit⁻¹ :) : R) + simp [smul_smul] /-- We can regard `coe_iSup_of_chain` as the statement that `(↑) : (Submodule R M) → Set M` is Scott continuous for the ω-complete partial order induced by the complete lattice structures. -/ diff --git a/Mathlib/LinearAlgebra/TensorProduct/Basic.lean b/Mathlib/LinearAlgebra/TensorProduct/Basic.lean index 75d5ed0bc25c0..64f12cf756387 100644 --- a/Mathlib/LinearAlgebra/TensorProduct/Basic.lean +++ b/Mathlib/LinearAlgebra/TensorProduct/Basic.lean @@ -1529,7 +1529,7 @@ instance CompatibleSMul.int : CompatibleSMul R ℤ M N := instance CompatibleSMul.unit {S} [Monoid S] [DistribMulAction S M] [DistribMulAction S N] [CompatibleSMul R S M N] : CompatibleSMul R Sˣ M N := - ⟨fun s m n => (CompatibleSMul.smul_tmul (s : S) m n : _)⟩ + ⟨fun s m n => CompatibleSMul.smul_tmul (s : S) m n⟩ end TensorProduct diff --git a/Mathlib/LinearAlgebra/TensorProduct/Graded/Internal.lean b/Mathlib/LinearAlgebra/TensorProduct/Graded/Internal.lean index c06613d1d7037..5c48e5305772b 100644 --- a/Mathlib/LinearAlgebra/TensorProduct/Graded/Internal.lean +++ b/Mathlib/LinearAlgebra/TensorProduct/Graded/Internal.lean @@ -252,7 +252,7 @@ instance instAlgebra : Algebra R (𝒜 ᵍ⊗[R] ℬ) where simp_rw [DirectSum.decompose_algebraMap, DirectSum.decompose_one, algebraMap_gradedMul] -- Qualified `map_smul` to avoid a TC timeout https://github.com/leanprover-community/mathlib4/pull/8386 erw [LinearMap.map_smul] - erw [LinearEquiv.symm_apply_apply] + simp lemma algebraMap_def (r : R) : algebraMap R (𝒜 ᵍ⊗[R] ℬ) r = algebraMap R A r ᵍ⊗ₜ[R] 1 := rfl diff --git a/Mathlib/LinearAlgebra/TensorProduct/Quotient.lean b/Mathlib/LinearAlgebra/TensorProduct/Quotient.lean index 21ad44e035faf..0c8559a3153a9 100644 --- a/Mathlib/LinearAlgebra/TensorProduct/Quotient.lean +++ b/Mathlib/LinearAlgebra/TensorProduct/Quotient.lean @@ -107,7 +107,7 @@ noncomputable def quotientTensorEquiv (m : Submodule R M) : erw [Submodule.map_id, Submodule.map_id] simp only [sup_eq_left] rw [map_range_eq_span_tmul, map_range_eq_span_tmul] - aesop) + simp) @[simp] lemma quotientTensorEquiv_apply_tmul_mk (m : Submodule R M) (x : M) (y : N) : @@ -137,7 +137,7 @@ noncomputable def tensorQuotientEquiv (n : Submodule R N) : erw [Submodule.map_id, Submodule.map_id] simp only [sup_eq_right] rw [map_range_eq_span_tmul, map_range_eq_span_tmul] - aesop) + simp) @[simp] lemma tensorQuotientEquiv_apply_mk_tmul (n : Submodule R N) (x : M) (y : N) : diff --git a/Mathlib/LinearAlgebra/TensorProduct/Tower.lean b/Mathlib/LinearAlgebra/TensorProduct/Tower.lean index b55604d42a282..33f45ceefbacc 100644 --- a/Mathlib/LinearAlgebra/TensorProduct/Tower.lean +++ b/Mathlib/LinearAlgebra/TensorProduct/Tower.lean @@ -94,7 +94,7 @@ See note [partially-applied ext lemmas]. -/ nonrec theorem curry_injective : Function.Injective (curry : (M ⊗ N →ₗ[A] P) → M →ₗ[A] N →ₗ[R] P) := fun _ _ h => LinearMap.restrictScalars_injective R <| - curry_injective <| (congr_arg (LinearMap.restrictScalars R) h : _) + curry_injective <| (congr_arg (LinearMap.restrictScalars R) h :) theorem ext {g h : M ⊗[R] N →ₗ[A] P} (H : ∀ x y, g (x ⊗ₜ y) = h (x ⊗ₜ y)) : g = h := curry_injective <| LinearMap.ext₂ H diff --git a/Mathlib/Logic/Basic.lean b/Mathlib/Logic/Basic.lean index 409b30a71b81e..56e15a7a87dc5 100644 --- a/Mathlib/Logic/Basic.lean +++ b/Mathlib/Logic/Basic.lean @@ -375,8 +375,6 @@ theorem forall_mem_comm {α β} [Membership α β] {s : β} {p : α → α → P (∀ a (_ : a ∈ s) b (_ : b ∈ s), p a b) ↔ ∀ a b, a ∈ s → b ∈ s → p a b := forall_cond_comm -@[deprecated (since := "2024-03-23")] alias ball_cond_comm := forall_cond_comm -@[deprecated (since := "2024-03-23")] alias ball_mem_comm := forall_mem_comm lemma ne_of_eq_of_ne {α : Sort*} {a b c : α} (h₁ : a = b) (h₂ : b ≠ c) : a ≠ c := h₁.symm ▸ h₂ lemma ne_of_ne_of_eq {α : Sort*} {a b c : α} (h₁ : a ≠ b) (h₂ : b = c) : a ≠ c := h₂ ▸ h₁ @@ -783,21 +781,12 @@ theorem BAll.imp_left (H : ∀ x, p x → q x) (h₁ : ∀ x, q x → r x) (x) ( theorem BEx.imp_left (H : ∀ x, p x → q x) : (∃ (x : _) (_ : p x), r x) → ∃ (x : _) (_ : q x), r x | ⟨x, hp, hr⟩ => ⟨x, H _ hp, hr⟩ -@[deprecated id (since := "2024-03-23")] -theorem ball_of_forall (h : ∀ x, p x) (x) : p x := h x - -@[deprecated forall_imp (since := "2024-03-23")] -theorem forall_of_ball (H : ∀ x, p x) (h : ∀ x, p x → q x) (x) : q x := h x <| H x - theorem exists_mem_of_exists (H : ∀ x, p x) : (∃ x, q x) → ∃ (x : _) (_ : p x), q x | ⟨x, hq⟩ => ⟨x, H x, hq⟩ theorem exists_of_exists_mem : (∃ (x : _) (_ : p x), q x) → ∃ x, q x | ⟨x, _, hq⟩ => ⟨x, hq⟩ -@[deprecated (since := "2024-03-23")] alias bex_of_exists := exists_mem_of_exists -@[deprecated (since := "2024-03-23")] alias exists_of_bex := exists_of_exists_mem -@[deprecated (since := "2024-03-23")] alias bex_imp := exists₂_imp theorem not_exists_mem : (¬∃ x h, P x h) ↔ ∀ x h, ¬P x h := exists₂_imp diff --git a/Mathlib/Logic/Embedding/Set.lean b/Mathlib/Logic/Embedding/Set.lean index 0c0ce8045a169..1c85ef2c28d32 100644 --- a/Mathlib/Logic/Embedding/Set.lean +++ b/Mathlib/Logic/Embedding/Set.lean @@ -160,6 +160,8 @@ variable {α ι : Type*} {s t r : Set α} range (Function.Embedding.sumSet h) = s ∪ t := by simp [Set.ext_iff] +open scoped Function -- required for scoped `on` notation + /-- For an indexed family `s : ι → Set α` of disjoint sets, the natural injection from the sigma-type `(i : ι) × ↑(s i)` to `α`. -/ @[simps] def Function.Embedding.sigmaSet {s : ι → Set α} (h : Pairwise (Disjoint on s)) : diff --git a/Mathlib/Logic/Encodable/Lattice.lean b/Mathlib/Logic/Encodable/Lattice.lean index 039681749b576..5ca2ec3f7aecc 100644 --- a/Mathlib/Logic/Encodable/Lattice.lean +++ b/Mathlib/Logic/Encodable/Lattice.lean @@ -44,6 +44,7 @@ theorem iUnion_decode₂_cases {f : β → Set α} {C : Set α → Prop} (H0 : C convert H1 b simp [Set.ext_iff] +open scoped Function in -- required for scoped `on` notation theorem iUnion_decode₂_disjoint_on {f : β → Set α} (hd : Pairwise (Disjoint on f)) : Pairwise (Disjoint on fun i => ⋃ b ∈ decode₂ β i, f b) := by rintro i j ij diff --git a/Mathlib/Logic/Equiv/Pairwise.lean b/Mathlib/Logic/Equiv/Pairwise.lean index a21f1515ce92a..ebf8a8f6c4853 100644 --- a/Mathlib/Logic/Equiv/Pairwise.lean +++ b/Mathlib/Logic/Equiv/Pairwise.lean @@ -10,6 +10,8 @@ import Mathlib.Logic.Pairwise # Interaction of equivalences with `Pairwise` -/ +open scoped Function -- required for scoped `on` notation + lemma EmbeddingLike.pairwise_comp {X : Type*} {Y : Type*} {F} [FunLike F Y X] [EmbeddingLike F Y X] (f : F) {p : X → X → Prop} (h : Pairwise p) : Pairwise (p on f) := h.comp_of_injective <| EmbeddingLike.injective f diff --git a/Mathlib/Logic/Function/Basic.lean b/Mathlib/Logic/Function/Basic.lean index a59d0b66dcbae..1eee4fc7e8759 100644 --- a/Mathlib/Logic/Function/Basic.lean +++ b/Mathlib/Logic/Function/Basic.lean @@ -844,13 +844,13 @@ protected theorem uncurry {α β γ : Type*} {f : α → β → γ} (hf : Inject /-- As a map from the left argument to a unary function, `f` is injective. -/ theorem left' (hf : Injective2 f) [Nonempty β] : Function.Injective f := fun _ _ h ↦ let ⟨b⟩ := ‹Nonempty β› - hf.left b <| (congr_fun h b : _) + hf.left b <| (congr_fun h b :) /-- As a map from the right argument to a unary function, `f` is injective. -/ theorem right' (hf : Injective2 f) [Nonempty α] : Function.Injective fun b a ↦ f a b := fun _ _ h ↦ let ⟨a⟩ := ‹Nonempty α› - hf.right a <| (congr_fun h a : _) + hf.right a <| (congr_fun h a :) theorem eq_iff (hf : Injective2 f) {a₁ a₂ b₁ b₂} : f a₁ b₁ = f a₂ b₂ ↔ a₁ = a₂ ∧ b₁ = b₂ := ⟨fun h ↦ hf h, fun ⟨h1, h2⟩ ↦ congr_arg₂ f h1 h2⟩ diff --git a/Mathlib/Logic/Function/Defs.lean b/Mathlib/Logic/Function/Defs.lean index 21054aab26345..0caf5f335d524 100644 --- a/Mathlib/Logic/Function/Defs.lean +++ b/Mathlib/Logic/Function/Defs.lean @@ -44,7 +44,7 @@ from `β` to `α`. -/ abbrev onFun (f : β → β → φ) (g : α → β) : α → α → φ := fun x y => f (g x) (g y) @[inherit_doc onFun] -infixl:2 " on " => onFun +scoped infixl:2 " on " => onFun abbrev swap {φ : α → β → Sort u₃} (f : ∀ x y, φ x y) : ∀ y x, φ x y := fun y x => f x y diff --git a/Mathlib/Logic/Godel/GodelBetaFunction.lean b/Mathlib/Logic/Godel/GodelBetaFunction.lean index 22435ce394cd4..092137dd0dea8 100644 --- a/Mathlib/Logic/Godel/GodelBetaFunction.lean +++ b/Mathlib/Logic/Godel/GodelBetaFunction.lean @@ -68,6 +68,7 @@ lemma coprimes_lt (a : Fin m → ℕ) (i) : a i < coprimes a i := by (le_add_right _ _)) simpa only [coprimes] using lt_of_lt_of_le h₁ h₂ +open scoped Function in -- required for scoped `on` notation private lemma pairwise_coprime_coprimes (a : Fin m → ℕ) : Pairwise (Coprime on coprimes a) := by intro i j hij wlog ltij : i < j diff --git a/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean b/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean index 3648b47ac5b12..8767171c0332e 100644 --- a/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean +++ b/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean @@ -222,6 +222,7 @@ theorem IsOpen.measurableSet (h : IsOpen s) : MeasurableSet s := theorem IsOpen.nullMeasurableSet {μ} (h : IsOpen s) : NullMeasurableSet s μ := h.measurableSet.nullMeasurableSet +open scoped Function in -- required for scoped `on` notation @[elab_as_elim] theorem MeasurableSet.induction_on_open {C : ∀ s : Set γ, MeasurableSet s → Prop} (isOpen : ∀ U (hU : IsOpen U), C U hU.measurableSet) diff --git a/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean b/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean index adcc5cb835bb0..7af987e92ef57 100644 --- a/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean +++ b/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean @@ -195,9 +195,6 @@ theorem aemeasurable_coe_nnreal_real_iff {f : α → ℝ≥0} {μ : Measure α} AEMeasurable (fun x => f x : α → ℝ) μ ↔ AEMeasurable f μ := ⟨fun h ↦ by simpa only [Real.toNNReal_coe] using h.real_toNNReal, AEMeasurable.coe_nnreal_real⟩ -@[deprecated (since := "2024-03-02")] -alias aEMeasurable_coe_nnreal_real_iff := aemeasurable_coe_nnreal_real_iff - /-- The set of finite `ℝ≥0∞` numbers is `MeasurableEquiv` to `ℝ≥0`. -/ def MeasurableEquiv.ennrealEquivNNReal : { r : ℝ≥0∞ | r ≠ ∞ } ≃ᵐ ℝ≥0 := ENNReal.neTopHomeomorphNNReal.toMeasurableEquiv @@ -283,17 +280,11 @@ theorem measurable_of_tendsto' {ι : Type*} {f : ι → α → ℝ≥0∞} {g : show Measurable fun y => liminf (fun n => (f (x n) y : ℝ≥0∞)) atTop exact .liminf fun n => hf (x n) -@[deprecated (since := "2024-03-09")] alias -_root_.measurable_of_tendsto_ennreal' := ENNReal.measurable_of_tendsto' - /-- A sequential limit of measurable `ℝ≥0∞` valued functions is measurable. -/ theorem measurable_of_tendsto {f : ℕ → α → ℝ≥0∞} {g : α → ℝ≥0∞} (hf : ∀ i, Measurable (f i)) (lim : Tendsto f atTop (𝓝 g)) : Measurable g := measurable_of_tendsto' atTop hf lim -@[deprecated (since := "2024-03-09")] alias -_root_.measurable_of_tendsto_ennreal := ENNReal.measurable_of_tendsto - /-- A limit (over a general filter) of a.e.-measurable `ℝ≥0∞` valued functions is a.e.-measurable. -/ lemma aemeasurable_of_tendsto' {ι : Type*} {f : ι → α → ℝ≥0∞} {g : α → ℝ≥0∞} @@ -453,17 +444,11 @@ theorem measurable_of_tendsto' {ι} {f : ι → α → ℝ≥0} {g : α → ℝ rw [tendsto_pi_nhds] at lim ⊢ exact fun x => (ENNReal.continuous_coe.tendsto (g x)).comp (lim x) -@[deprecated (since := "2024-03-09")] alias -_root_.measurable_of_tendsto_nnreal' := NNReal.measurable_of_tendsto' - /-- A sequential limit of measurable `ℝ≥0` valued functions is measurable. -/ theorem measurable_of_tendsto {f : ℕ → α → ℝ≥0} {g : α → ℝ≥0} (hf : ∀ i, Measurable (f i)) (lim : Tendsto f atTop (𝓝 g)) : Measurable g := measurable_of_tendsto' atTop hf lim -@[deprecated (since := "2024-03-09")] alias -_root_.measurable_of_tendsto_nnreal := NNReal.measurable_of_tendsto - end NNReal namespace EReal diff --git a/Mathlib/MeasureTheory/Constructions/Pi.lean b/Mathlib/MeasureTheory/Constructions/Pi.lean index 963c8dcd5082f..102c61a082662 100644 --- a/Mathlib/MeasureTheory/Constructions/Pi.lean +++ b/Mathlib/MeasureTheory/Constructions/Pi.lean @@ -4,11 +4,12 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn -/ import Mathlib.MeasureTheory.Group.Measure +import Mathlib.MeasureTheory.MeasurableSpace.Pi import Mathlib.MeasureTheory.Measure.Prod import Mathlib.Topology.Constructions /-! -# Product measures +# Indexed product measures In this file we define and prove properties about finite products of measures (and at some point, countable products of measures). @@ -49,7 +50,6 @@ finitary product measure -/ - noncomputable section open Function Set MeasureTheory.OuterMeasure Filter MeasurableSpace Encodable @@ -60,88 +60,6 @@ universe u v variable {ι ι' : Type*} {α : ι → Type*} -/-! We start with some measurability properties -/ - -lemma MeasurableSpace.pi_eq_generateFrom_projections {mα : ∀ i, MeasurableSpace (α i)} : - pi = generateFrom {B | ∃ (i : ι) (A : Set (α i)), MeasurableSet A ∧ eval i ⁻¹' A = B} := by - simp only [pi, ← generateFrom_iUnion_measurableSet, iUnion_setOf, measurableSet_comap] - -/-- Boxes formed by π-systems form a π-system. -/ -theorem IsPiSystem.pi {C : ∀ i, Set (Set (α i))} (hC : ∀ i, IsPiSystem (C i)) : - IsPiSystem (pi univ '' pi univ C) := by - rintro _ ⟨s₁, hs₁, rfl⟩ _ ⟨s₂, hs₂, rfl⟩ hst - rw [← pi_inter_distrib] at hst ⊢; rw [univ_pi_nonempty_iff] at hst - exact mem_image_of_mem _ fun i _ => hC i _ (hs₁ i (mem_univ i)) _ (hs₂ i (mem_univ i)) (hst i) - -/-- Boxes form a π-system. -/ -theorem isPiSystem_pi [∀ i, MeasurableSpace (α i)] : - IsPiSystem (pi univ '' pi univ fun i => { s : Set (α i) | MeasurableSet s }) := - IsPiSystem.pi fun _ => isPiSystem_measurableSet - -section Finite - -variable [Finite ι] - -/-- Boxes of countably spanning sets are countably spanning. -/ -theorem IsCountablySpanning.pi {C : ∀ i, Set (Set (α i))} (hC : ∀ i, IsCountablySpanning (C i)) : - IsCountablySpanning (pi univ '' pi univ C) := by - choose s h1s h2s using hC - cases nonempty_encodable (ι → ℕ) - let e : ℕ → ι → ℕ := fun n => (@decode (ι → ℕ) _ n).iget - refine ⟨fun n => Set.pi univ fun i => s i (e n i), fun n => - mem_image_of_mem _ fun i _ => h1s i _, ?_⟩ - simp_rw - [e, (surjective_decode_iget (ι → ℕ)).iUnion_comp fun x => Set.pi univ fun i => s i (x i), - iUnion_univ_pi s, h2s, pi_univ] - -/-- The product of generated σ-algebras is the one generated by boxes, if both generating sets - are countably spanning. -/ -theorem generateFrom_pi_eq {C : ∀ i, Set (Set (α i))} (hC : ∀ i, IsCountablySpanning (C i)) : - (@MeasurableSpace.pi _ _ fun i => generateFrom (C i)) = - generateFrom (pi univ '' pi univ C) := by - cases nonempty_encodable ι - apply le_antisymm - · refine iSup_le ?_; intro i; rw [comap_generateFrom] - apply generateFrom_le; rintro _ ⟨s, hs, rfl⟩ - choose t h1t h2t using hC - simp_rw [eval_preimage, ← h2t] - rw [← @iUnion_const _ ℕ _ s] - have : Set.pi univ (update (fun i' : ι => iUnion (t i')) i (⋃ _ : ℕ, s)) = - Set.pi univ fun k => ⋃ j : ℕ, - @update ι (fun i' => Set (α i')) _ (fun i' => t i' j) i s k := by - ext; simp_rw [mem_univ_pi]; apply forall_congr'; intro i' - by_cases h : i' = i - · subst h; simp - · rw [← Ne] at h; simp [h] - rw [this, ← iUnion_univ_pi] - apply MeasurableSet.iUnion - intro n; apply measurableSet_generateFrom - apply mem_image_of_mem; intro j _; dsimp only - by_cases h : j = i - · subst h; rwa [update_self] - · rw [update_of_ne h]; apply h1t - · apply generateFrom_le; rintro _ ⟨s, hs, rfl⟩ - rw [univ_pi_eq_iInter]; apply MeasurableSet.iInter; intro i - apply @measurable_pi_apply _ _ (fun i => generateFrom (C i)) - exact measurableSet_generateFrom (hs i (mem_univ i)) - -/-- If `C` and `D` generate the σ-algebras on `α` resp. `β`, then rectangles formed by `C` and `D` - generate the σ-algebra on `α × β`. -/ -theorem generateFrom_eq_pi [h : ∀ i, MeasurableSpace (α i)] {C : ∀ i, Set (Set (α i))} - (hC : ∀ i, generateFrom (C i) = h i) (h2C : ∀ i, IsCountablySpanning (C i)) : - generateFrom (pi univ '' pi univ C) = MeasurableSpace.pi := by - simp only [← funext hC, generateFrom_pi_eq h2C] - -/-- The product σ-algebra is generated from boxes, i.e. `s ×ˢ t` for sets `s : set α` and - `t : set β`. -/ -theorem generateFrom_pi [∀ i, MeasurableSpace (α i)] : - generateFrom (pi univ '' pi univ fun i => { s : Set (α i) | MeasurableSet s }) = - MeasurableSpace.pi := - generateFrom_eq_pi (fun _ => generateFrom_measurableSet) fun _ => - isCountablySpanning_measurableSet - -end Finite - namespace MeasureTheory variable [Fintype ι] {m : ∀ i, OuterMeasure (α i)} diff --git a/Mathlib/MeasureTheory/Constructions/Polish/Basic.lean b/Mathlib/MeasureTheory/Constructions/Polish/Basic.lean index ff610fee246c5..b016d6eba33b4 100644 --- a/Mathlib/MeasureTheory/Constructions/Polish/Basic.lean +++ b/Mathlib/MeasureTheory/Constructions/Polish/Basic.lean @@ -251,7 +251,7 @@ theorem AnalyticSet.iInter [hι : Nonempty ι] [Countable ι] [T2Space α] {s : apply Subset.antisymm · rintro y ⟨x, rfl⟩ refine mem_iInter.2 fun n => ?_ - have : f n ((x : γ) n) = F x := (mem_iInter.1 x.2 n : _) + have : f n ((x : γ) n) = F x := (mem_iInter.1 x.2 n :) rw [← this, ← f_range n] exact mem_range_self _ · intro y hy @@ -740,10 +740,10 @@ theorem measurableSet_range_of_continuous_injective {β : Type*} [TopologicalSpa 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 : _) + (mem_iInter.1 this h :) have B : x ∈ q ⟨(s n, s m), h.symm⟩ \ q ⟨(s m, s n), h⟩ := haveI := mem_iInter.1 (hxs n).2 (s m) - (mem_iInter.1 this h.symm : _) + (mem_iInter.1 this h.symm :) exact A.2 B.1 -- the points `y n` are nearby, and therefore they form a Cauchy sequence. have cauchy_y : CauchySeq y := by diff --git a/Mathlib/MeasureTheory/Covering/Besicovitch.lean b/Mathlib/MeasureTheory/Covering/Besicovitch.lean index a21afb385c697..9f8b24ead46f2 100644 --- a/Mathlib/MeasureTheory/Covering/Besicovitch.lean +++ b/Mathlib/MeasureTheory/Covering/Besicovitch.lean @@ -791,14 +791,17 @@ theorem exists_disjoint_closedBall_covering_ae_of_finiteMeasure_aux (μ : Measur rw [← Nat.succ_eq_add_one, u_succ] exact (hF (u n) (Pu n)).1 -/-- The measurable Besicovitch covering theorem. Assume that, for any `x` in a set `s`, -one is given a set of admissible closed balls centered at `x`, with arbitrarily small radii. -Then there exists a disjoint covering of almost all `s` by admissible closed balls centered at some -points of `s`. -This version requires that the underlying measure is sigma-finite, and that the space has the +/-- The measurable **Besicovitch covering theorem**. + +Assume that, for any `x` in a set `s`, one is given a set of admissible closed balls centered at +`x`, with arbitrarily small radii. Then there exists a disjoint covering of almost all `s` by +admissible closed balls centered at some points of `s`. + +This version requires the underlying measure to be sigma-finite, and the space to have the Besicovitch covering property (which is satisfied for instance by normed real vector spaces). It expresses the conclusion in a slightly awkward form (with a subset of `α × ℝ`) coming from the proof technique. + For a version giving the conclusion in a nicer form, see `exists_disjoint_closedBall_covering_ae`. -/ theorem exists_disjoint_closedBall_covering_ae_aux (μ : Measure α) [SFinite μ] (f : α → Set ℝ) @@ -813,12 +816,14 @@ theorem exists_disjoint_closedBall_covering_ae_aux (μ : Measure α) [SFinite μ ⟨t, t_count, ts, tr, tν, tdisj⟩ exact ⟨t, t_count, ts, tr, hμν tν, tdisj⟩ -/-- The measurable Besicovitch covering theorem. Assume that, for any `x` in a set `s`, -one is given a set of admissible closed balls centered at `x`, with arbitrarily small radii. -Then there exists a disjoint covering of almost all `s` by admissible closed balls centered at some -points of `s`. We can even require that the radius at `x` is bounded by a given function `R x`. -(Take `R = 1` if you don't need this additional feature). -This version requires that the underlying measure is sigma-finite, and that the space has the +/-- The measurable **Besicovitch covering theorem**. + +Assume that, for any `x` in a set `s`, one is given a set of admissible closed balls centered at +`x`, with arbitrarily small radii. Then there exists a disjoint covering of almost all `s` by +admissible closed balls centered at some points of `s`. We can even require that the radius at `x` +is bounded by a given function `R x`. (Take `R = 1` if you don't need this additional feature). + +This version requires the underlying measure to be sigma-finite, and the space to have the Besicovitch covering property (which is satisfied for instance by normed real vector spaces). -/ theorem exists_disjoint_closedBall_covering_ae (μ : Measure α) [SFinite μ] (f : α → Set ℝ) diff --git a/Mathlib/MeasureTheory/Covering/BesicovitchVectorSpace.lean b/Mathlib/MeasureTheory/Covering/BesicovitchVectorSpace.lean index e0b6251d52956..33e6f5d60ba87 100644 --- a/Mathlib/MeasureTheory/Covering/BesicovitchVectorSpace.lean +++ b/Mathlib/MeasureTheory/Covering/BesicovitchVectorSpace.lean @@ -99,6 +99,7 @@ section variable [NormedSpace ℝ E] [FiniteDimensional ℝ E] +open scoped Function in -- required for scoped `on` notation /-- Any `1`-separated set in the ball of radius `2` has cardinality at most `5 ^ dim`. This is useful to show that the supremum in the definition of `Besicovitch.multiplicity E` is well behaved. -/ diff --git a/Mathlib/MeasureTheory/Covering/Vitali.lean b/Mathlib/MeasureTheory/Covering/Vitali.lean index 416a61a4c48c3..5b2934b6db6d8 100644 --- a/Mathlib/MeasureTheory/Covering/Vitali.lean +++ b/Mathlib/MeasureTheory/Covering/Vitali.lean @@ -198,12 +198,14 @@ theorem exists_disjoint_subfamily_covering_enlargement_closedBall alias exists_disjoint_subfamily_covering_enlargment_closedBall := exists_disjoint_subfamily_covering_enlargement_closedBall -/-- The measurable Vitali covering theorem. Assume one is given a family `t` of closed sets with -nonempty interior, such that each `a ∈ t` is included in a ball `B (x, r)` and covers a definite -proportion of the ball `B (x, 3 r)` for a given measure `μ` (think of the situation where `μ` is -a doubling measure and `t` is a family of balls). Consider a (possibly non-measurable) set `s` -at which the family is fine, i.e., every point of `s` belongs to arbitrarily small elements of `t`. -Then one can extract from `t` a disjoint subfamily that covers almost all `s`. +/-- The measurable **Vitali covering theorem**. + +Assume one is given a family `t` of closed sets with nonempty interior, such that each `a ∈ t` is +included in a ball `B (x, r)` and covers a definite proportion of the ball `B (x, 3 r)` for a given +measure `μ` (think of the situation where `μ` is a doubling measure and `t` is a family of balls). +Consider a (possibly non-measurable) set `s` at which the family is fine, i.e., every point of `s` +belongs to arbitrarily small elements of `t`. Then one can extract from `t` a disjoint subfamily +that covers almost all `s`. For more flexibility, we give a statement with a parameterized family of sets. -/ diff --git a/Mathlib/MeasureTheory/Covering/VitaliFamily.lean b/Mathlib/MeasureTheory/Covering/VitaliFamily.lean index 0cfed115a59d7..7671daaa30640 100644 --- a/Mathlib/MeasureTheory/Covering/VitaliFamily.lean +++ b/Mathlib/MeasureTheory/Covering/VitaliFamily.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 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.MeasureTheory.Measure.MeasureSpace +import Mathlib.MeasureTheory.Measure.AbsolutelyContinuous /-! # Vitali families @@ -131,6 +131,7 @@ theorem index_subset : ∀ p : X × Set X, p ∈ h.index → p.1 ∈ s := theorem covering_disjoint : h.index.PairwiseDisjoint h.covering := h.exists_disjoint_covering_ae.choose_spec.2.1 +open scoped Function in -- required for scoped `on` notation theorem covering_disjoint_subtype : Pairwise (Disjoint on fun x : h.index => h.covering x) := (pairwise_subtype_iff_pairwise_set _ _).2 h.covering_disjoint diff --git a/Mathlib/MeasureTheory/Decomposition/SignedHahn.lean b/Mathlib/MeasureTheory/Decomposition/SignedHahn.lean index f6fe47d6c9795..0a690c953d507 100644 --- a/Mathlib/MeasureTheory/Decomposition/SignedHahn.lean +++ b/Mathlib/MeasureTheory/Decomposition/SignedHahn.lean @@ -211,6 +211,7 @@ private theorem restrictNonposSeq_disjoint' {n m : ℕ} (h : n < m) : (someExistsOneDivLT_subset hx₂).2 (Set.mem_iUnion.2 ⟨n, Set.mem_iUnion.2 ⟨Nat.lt_succ_iff.mp h, hx₁⟩⟩) +open scoped Function in -- required for scoped `on` notation private theorem restrictNonposSeq_disjoint : Pairwise (Disjoint on restrictNonposSeq s i) := by intro n m h rw [Function.onFun, Set.disjoint_iff_inter_eq_empty] diff --git a/Mathlib/MeasureTheory/Function/Jacobian.lean b/Mathlib/MeasureTheory/Function/Jacobian.lean index f4914c6260a11..bfb9326f21a71 100644 --- a/Mathlib/MeasureTheory/Function/Jacobian.lean +++ b/Mathlib/MeasureTheory/Function/Jacobian.lean @@ -242,6 +242,8 @@ theorem exists_closed_cover_approximatesLinearOn_of_hasFDerivWithinAt [SecondCou variable [MeasurableSpace E] [BorelSpace E] (μ : Measure E) [IsAddHaarMeasure μ] +open scoped Function -- required for scoped `on` notation + /-- Assume that a function `f` has a derivative at every point of a set `s`. Then one may partition `s` into countably many disjoint relatively measurable sets (i.e., intersections of `s` with measurable sets `t n`) on which `f` is well approximated by linear maps `A n`. -/ diff --git a/Mathlib/MeasureTheory/Function/L1Space.lean b/Mathlib/MeasureTheory/Function/L1Space.lean index 30d728e7e06fc..0005ef1d22c72 100644 --- a/Mathlib/MeasureTheory/Function/L1Space.lean +++ b/Mathlib/MeasureTheory/Function/L1Space.lean @@ -173,9 +173,6 @@ theorem HasFiniteIntegral.of_finite [Finite α] [IsFiniteMeasure μ] {f : α → let ⟨_⟩ := nonempty_fintype α hasFiniteIntegral_of_bounded <| ae_of_all μ <| norm_le_pi_norm f -@[deprecated (since := "2024-02-05")] -alias hasFiniteIntegral_of_fintype := HasFiniteIntegral.of_finite - theorem HasFiniteIntegral.mono_measure {f : α → β} (h : HasFiniteIntegral f ν) (hμ : μ ≤ ν) : HasFiniteIntegral f μ := lt_of_le_of_lt (lintegral_mono' hμ le_rfl) h @@ -474,7 +471,6 @@ lemma Integrable.of_finite [Finite α] [MeasurableSingletonClass α] [IsFiniteMe @[deprecated Integrable.of_finite (since := "2024-10-05"), nolint deprecatedNoSince] lemma Integrable.of_isEmpty [IsEmpty α] {f : α → β} : Integrable f μ := .of_finite -@[deprecated (since := "2024-02-05")] alias integrable_of_fintype := Integrable.of_finite theorem Memℒp.integrable_norm_rpow {f : α → β} {p : ℝ≥0∞} (hf : Memℒp f p μ) (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) : Integrable (fun x : α => ‖f x‖ ^ p.toReal) μ := by diff --git a/Mathlib/MeasureTheory/Function/LpSeminorm/CompareExp.lean b/Mathlib/MeasureTheory/Function/LpSeminorm/CompareExp.lean index 1df002f2d7299..6c1499cfe95ca 100644 --- a/Mathlib/MeasureTheory/Function/LpSeminorm/CompareExp.lean +++ b/Mathlib/MeasureTheory/Function/LpSeminorm/CompareExp.lean @@ -306,7 +306,7 @@ variable {𝕜 α E F : Type*} {m : MeasurableSpace α} {μ : Measure α} [Norme theorem eLpNorm_smul_le_eLpNorm_top_mul_eLpNorm (p : ℝ≥0∞) (hf : AEStronglyMeasurable f μ) (φ : α → 𝕜) : eLpNorm (φ • f) p μ ≤ eLpNorm φ ∞ μ * eLpNorm f p μ := (eLpNorm_le_eLpNorm_top_mul_eLpNorm p φ hf (· • ·) - (Eventually.of_forall fun _ => nnnorm_smul_le _ _) : _) + (Eventually.of_forall fun _ => nnnorm_smul_le _ _) :) @[deprecated (since := "2024-07-27")] alias snorm_smul_le_snorm_top_mul_snorm := eLpNorm_smul_le_eLpNorm_top_mul_eLpNorm @@ -314,7 +314,7 @@ alias snorm_smul_le_snorm_top_mul_snorm := eLpNorm_smul_le_eLpNorm_top_mul_eLpNo theorem eLpNorm_smul_le_eLpNorm_mul_eLpNorm_top (p : ℝ≥0∞) (f : α → E) {φ : α → 𝕜} (hφ : AEStronglyMeasurable φ μ) : eLpNorm (φ • f) p μ ≤ eLpNorm φ p μ * eLpNorm f ∞ μ := (eLpNorm_le_eLpNorm_mul_eLpNorm_top p hφ f (· • ·) - (Eventually.of_forall fun _ => nnnorm_smul_le _ _) : _) + (Eventually.of_forall fun _ => nnnorm_smul_le _ _) :) @[deprecated (since := "2024-07-27")] alias snorm_smul_le_snorm_mul_snorm_top := eLpNorm_smul_le_eLpNorm_mul_eLpNorm_top diff --git a/Mathlib/MeasureTheory/Function/SimpleFunc.lean b/Mathlib/MeasureTheory/Function/SimpleFunc.lean index 05df7e2bf5776..3d030c368a939 100644 --- a/Mathlib/MeasureTheory/Function/SimpleFunc.lean +++ b/Mathlib/MeasureTheory/Function/SimpleFunc.lean @@ -77,7 +77,6 @@ def ofFinite [Finite α] [MeasurableSingletonClass α] (f : α → β) : α → measurableSet_fiber' x := (toFinite (f ⁻¹' {x})).measurableSet finite_range' := Set.finite_range f -@[deprecated (since := "2024-02-05")] alias ofFintype := ofFinite /-- Simple function defined on the empty type. -/ def ofIsEmpty [IsEmpty α] : α →ₛ β := ofFinite isEmptyElim diff --git a/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean b/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean index b8cf7ba03d100..d64f4c733873a 100644 --- a/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean +++ b/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean @@ -469,7 +469,7 @@ attribute [local instance] simpleFunc.module /-- If `E` is a normed space, `Lp.simpleFunc E p μ` is a normed space. Not declared as an instance as it is (as of writing) used only in the construction of the Bochner integral. -/ protected theorem boundedSMul [Fact (1 ≤ p)] : BoundedSMul 𝕜 (Lp.simpleFunc E p μ) := - BoundedSMul.of_norm_smul_le fun r f => (norm_smul_le r (f : Lp E p μ) : _) + BoundedSMul.of_norm_smul_le fun r f => (norm_smul_le r (f : Lp E p μ) :) attribute [local instance] simpleFunc.boundedSMul diff --git a/Mathlib/MeasureTheory/Function/SpecialFunctions/Basic.lean b/Mathlib/MeasureTheory/Function/SpecialFunctions/Basic.lean index e072184cd1af6..2074623647f98 100644 --- a/Mathlib/MeasureTheory/Function/SpecialFunctions/Basic.lean +++ b/Mathlib/MeasureTheory/Function/SpecialFunctions/Basic.lean @@ -28,6 +28,8 @@ open NNReal ENNReal MeasureTheory namespace Real +variable {α : Type*} {_ : MeasurableSpace α} {f : α → ℝ} {μ : MeasureTheory.Measure α} + @[measurability] theorem measurable_exp : Measurable exp := continuous_exp.measurable @@ -37,20 +39,24 @@ theorem measurable_log : Measurable log := measurable_of_measurable_on_compl_singleton 0 <| Continuous.measurable <| continuousOn_iff_continuous_restrict.1 continuousOn_log -lemma measurable_of_measurable_exp {α : Type*} {_ : MeasurableSpace α} {f : α → ℝ} - (hf : Measurable (fun x ↦ exp (f x))) : +lemma measurable_of_measurable_exp (hf : Measurable (fun x ↦ exp (f x))) : Measurable f := by have : f = fun x ↦ log (exp (f x)) := by ext; rw [log_exp] rw [this] exact measurable_log.comp hf -lemma aemeasurable_of_aemeasurable_exp {α : Type*} {_ : MeasurableSpace α} {f : α → ℝ} - {μ : MeasureTheory.Measure α} (hf : AEMeasurable (fun x ↦ exp (f x)) μ) : +lemma aemeasurable_of_aemeasurable_exp (hf : AEMeasurable (fun x ↦ exp (f x)) μ) : AEMeasurable f μ := by have : f = fun x ↦ log (exp (f x)) := by ext; rw [log_exp] rw [this] exact measurable_log.comp_aemeasurable hf +lemma aemeasurable_of_aemeasurable_exp_mul {t : ℝ} + (ht : t ≠ 0) (hf : AEMeasurable (fun x ↦ exp (t * f x)) μ) : + AEMeasurable f μ := by + simpa only [mul_div_cancel_left₀ _ ht] + using (aemeasurable_of_aemeasurable_exp hf).div (aemeasurable_const (b := t)) + @[measurability] theorem measurable_sin : Measurable sin := continuous_sin.measurable diff --git a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean index 757c23abb972f..3fcd3f10b57d4 100644 --- a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean +++ b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean @@ -130,14 +130,6 @@ theorem StronglyMeasurable.of_finite [Finite α] {_ : MeasurableSpace α} {f : α → β} : StronglyMeasurable f := ⟨fun _ => SimpleFunc.ofFinite f, fun _ => tendsto_const_nhds⟩ -@[deprecated (since := "2024-02-05")] -alias stronglyMeasurable_of_fintype := StronglyMeasurable.of_finite - -@[deprecated StronglyMeasurable.of_finite (since := "2024-02-06")] -theorem stronglyMeasurable_of_isEmpty [IsEmpty α] {_ : MeasurableSpace α} [TopologicalSpace β] - (f : α → β) : StronglyMeasurable f := - .of_finite - theorem stronglyMeasurable_const {α β} {_ : MeasurableSpace α} [TopologicalSpace β] {b : β} : StronglyMeasurable fun _ : α => b := ⟨fun _ => SimpleFunc.const α b, fun _ => tendsto_const_nhds⟩ @@ -1162,7 +1154,6 @@ theorem mono_measure {ν : Measure α} (hf : AEStronglyMeasurable f μ) (h : ν protected lemma mono_ac (h : ν ≪ μ) (hμ : AEStronglyMeasurable f μ) : AEStronglyMeasurable f ν := let ⟨g, hg, hg'⟩ := hμ; ⟨g, hg, h.ae_eq hg'⟩ -@[deprecated (since := "2024-02-15")] protected alias mono' := AEStronglyMeasurable.mono_ac theorem mono_set {s t} (h : s ⊆ t) (ht : AEStronglyMeasurable f (μ.restrict t)) : AEStronglyMeasurable f (μ.restrict s) := diff --git a/Mathlib/MeasureTheory/Group/AddCircle.lean b/Mathlib/MeasureTheory/Group/AddCircle.lean index 2e2847f50467b..3f20417353a69 100644 --- a/Mathlib/MeasureTheory/Group/AddCircle.lean +++ b/Mathlib/MeasureTheory/Group/AddCircle.lean @@ -95,7 +95,7 @@ theorem volume_of_add_preimage_eq (s I : Set <| AddCircle T) (u x : AddCircle T) let G := AddSubgroup.zmultiples u haveI : Fintype G := @Fintype.ofFinite _ hu.finite_zmultiples.to_subtype 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 : _) + 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] diff --git a/Mathlib/MeasureTheory/Group/Arithmetic.lean b/Mathlib/MeasureTheory/Group/Arithmetic.lean index eaba1a41a370d..0d9b2426dd69b 100644 --- a/Mathlib/MeasureTheory/Group/Arithmetic.lean +++ b/Mathlib/MeasureTheory/Group/Arithmetic.lean @@ -662,7 +662,7 @@ instance Units.instMeasurableSpace : MeasurableSpace Mˣ := MeasurableSpace.coma @[to_additive] instance Units.measurableSMul : MeasurableSMul Mˣ β where - measurable_const_smul c := (measurable_const_smul (c : M) : _) + measurable_const_smul c := measurable_const_smul (c : M) measurable_smul_const x := (measurable_smul_const x : Measurable fun c : M => c • x).comp MeasurableSpace.le_map_comap diff --git a/Mathlib/MeasureTheory/Group/Defs.lean b/Mathlib/MeasureTheory/Group/Defs.lean index c68e87b9ad78f..d57a070570695 100644 --- a/Mathlib/MeasureTheory/Group/Defs.lean +++ b/Mathlib/MeasureTheory/Group/Defs.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Floris van Doorn. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn, Yury Kudryashov -/ -import Mathlib.MeasureTheory.Measure.MeasureSpace +import Mathlib.MeasureTheory.Measure.Map /-! # Definitions about invariant measures diff --git a/Mathlib/MeasureTheory/Integral/Bochner.lean b/Mathlib/MeasureTheory/Integral/Bochner.lean index dd20d7982e08b..2720940a760a3 100644 --- a/Mathlib/MeasureTheory/Integral/Bochner.lean +++ b/Mathlib/MeasureTheory/Integral/Bochner.lean @@ -1367,6 +1367,9 @@ theorem norm_integral_le_integral_norm (f : α → G) : ‖∫ a, f a ∂μ‖ · rw [integral_non_aestronglyMeasurable h, norm_zero] exact integral_nonneg_of_ae le_ae +lemma abs_integral_le_integral_abs {f : α → ℝ} : |∫ a, f a ∂μ| ≤ ∫ a, |f a| ∂μ := + norm_integral_le_integral_norm f + theorem norm_integral_le_of_norm_le {f : α → G} {g : α → ℝ} (hg : Integrable g μ) (h : ∀ᵐ x ∂μ, ‖f x‖ ≤ g x) : ‖∫ x, f x ∂μ‖ ≤ ∫ x, g x ∂μ := calc diff --git a/Mathlib/MeasureTheory/Integral/Lebesgue.lean b/Mathlib/MeasureTheory/Integral/Lebesgue.lean index 71735659713fe..b0cbd7bec0648 100644 --- a/Mathlib/MeasureTheory/Integral/Lebesgue.lean +++ b/Mathlib/MeasureTheory/Integral/Lebesgue.lean @@ -1304,6 +1304,8 @@ theorem lintegral_tsum [Countable β] {f : β → α → ℝ≥0∞} (hf : ∀ i open Measure +open scoped Function -- required for scoped `on` notation + theorem lintegral_iUnion₀ [Countable β] {s : β → Set α} (hm : ∀ i, NullMeasurableSet (s i) μ) (hd : Pairwise (AEDisjoint μ on s)) (f : α → ℝ≥0∞) : ∫⁻ a in ⋃ i, s i, f a ∂μ = ∑' i, ∫⁻ a in s i, f a ∂μ := by diff --git a/Mathlib/MeasureTheory/Integral/PeakFunction.lean b/Mathlib/MeasureTheory/Integral/PeakFunction.lean index 20b7988c88759..ea9c6d69cf046 100644 --- a/Mathlib/MeasureTheory/Integral/PeakFunction.lean +++ b/Mathlib/MeasureTheory/Integral/PeakFunction.lean @@ -83,10 +83,6 @@ theorem integrableOn_peak_smul_of_integrableOn_of_tendsto convert A.union B simp only [diff_union_inter] -@[deprecated (since := "2024-02-20")] -alias integrableOn_peak_smul_of_integrableOn_of_continuousWithinAt := - integrableOn_peak_smul_of_integrableOn_of_tendsto - /-- If a sequence of peak functions `φᵢ` converges uniformly to zero away from a point `x₀` and its integral on some finite-measure neighborhood of `x₀` converges to `1`, and `g` is integrable and has a limit `a` at `x₀`, then `∫ φᵢ • g` converges to `a`. @@ -176,10 +172,6 @@ theorem tendsto_setIntegral_peak_smul_of_integrableOn_of_tendsto_aux _ ≤ ‖∫ x in s \ u, φ i x • g x ∂μ‖ + ‖∫ x in s ∩ u, φ i x • g x ∂μ‖ := norm_add_le _ _ _ ≤ (δ * ∫ x in s, ‖g x‖ ∂μ) + 2 * δ := add_le_add C B -@[deprecated (since := "2024-02-20")] -alias tendsto_setIntegral_peak_smul_of_integrableOn_of_continuousWithinAt_aux := - tendsto_setIntegral_peak_smul_of_integrableOn_of_tendsto_aux - variable [CompleteSpace E] /-- If a sequence of peak functions `φᵢ` converges uniformly to zero away from a point `x₀` and its @@ -220,10 +212,6 @@ theorem tendsto_setIntegral_peak_smul_of_integrableOn_of_tendsto rw [restrict_restrict ht, inter_eq_left.mpr hts] exact .of_integral_ne_zero (fun h ↦ by simp [h] at h'i) -@[deprecated (since := "2024-02-20")] -alias tendsto_setIntegral_peak_smul_of_integrableOn_of_continuousWithinAt := - tendsto_setIntegral_peak_smul_of_integrableOn_of_tendsto - /-- If a sequence of peak functions `φᵢ` converges uniformly to zero away from a point `x₀` and its integral on some finite-measure neighborhood of `x₀` converges to `1`, and `g` is integrable and has a limit `a` at `x₀`, then `∫ φᵢ • g` converges to `a`. -/ diff --git a/Mathlib/MeasureTheory/Integral/Prod.lean b/Mathlib/MeasureTheory/Integral/Prod.lean index 7b77ffcd619e8..886d5b00bc910 100644 --- a/Mathlib/MeasureTheory/Integral/Prod.lean +++ b/Mathlib/MeasureTheory/Integral/Prod.lean @@ -301,7 +301,7 @@ theorem Integrable.integral_prod_left ⦃f : α × β → E⦄ (hf : Integrable (norm_integral_le_integral_norm _).trans_eq <| (norm_of_nonneg <| integral_nonneg_of_ae <| - Eventually.of_forall fun y => (norm_nonneg (f (x, y)) : _)).symm + Eventually.of_forall fun y => (norm_nonneg (f (x, y)) :)).symm theorem Integrable.integral_prod_right [SFinite μ] ⦃f : α × β → E⦄ (hf : Integrable f (μ.prod ν)) : Integrable (fun y => ∫ x, f (x, y) ∂μ) ν := diff --git a/Mathlib/MeasureTheory/Integral/SetToL1.lean b/Mathlib/MeasureTheory/Integral/SetToL1.lean index 7c6e7f8e1b323..9c3a4587da7c6 100644 --- a/Mathlib/MeasureTheory/Integral/SetToL1.lean +++ b/Mathlib/MeasureTheory/Integral/SetToL1.lean @@ -527,9 +527,6 @@ theorem norm_setToSimpleFunc_le_sum_opNorm {m : MeasurableSpace α} (T : Set α _ ≤ ∑ x ∈ f.range, ‖T (f ⁻¹' {x})‖ * ‖x‖ := by refine Finset.sum_le_sum fun b _ => ?_; simp_rw [ContinuousLinearMap.le_opNorm] -@[deprecated (since := "2024-02-02")] -alias norm_setToSimpleFunc_le_sum_op_norm := norm_setToSimpleFunc_le_sum_opNorm - theorem norm_setToSimpleFunc_le_sum_mul_norm (T : Set α → F →L[ℝ] F') {C : ℝ} (hT_norm : ∀ s, MeasurableSet s → ‖T s‖ ≤ C * (μ s).toReal) (f : α →ₛ F) : ‖f.setToSimpleFunc T‖ ≤ C * ∑ x ∈ f.range, (μ (f ⁻¹' {x})).toReal * ‖x‖ := diff --git a/Mathlib/MeasureTheory/MeasurableSpace/Embedding.lean b/Mathlib/MeasureTheory/MeasurableSpace/Embedding.lean index d597d44b6f446..b7600589ddd7c 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/Embedding.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/Embedding.lean @@ -428,8 +428,7 @@ def Set.rangeInl : (range Sum.inl : Set (α ⊕ β)) ≃ᵐ α where toEquiv := Equiv.Set.rangeInl α β measurable_toFun s (hs : MeasurableSet s) := by refine ⟨_, hs.inl_image, Set.ext ?_⟩ - rintro ⟨ab, a, rfl⟩ - simp [Set.range_inl] + simp measurable_invFun := Measurable.subtype_mk measurable_inl /-- `β` is equivalent to its image in `α ⊕ β` as measurable spaces. -/ @@ -437,8 +436,7 @@ def Set.rangeInr : (range Sum.inr : Set (α ⊕ β)) ≃ᵐ β where toEquiv := Equiv.Set.rangeInr α β measurable_toFun s (hs : MeasurableSet s) := by refine ⟨_, hs.inr_image, Set.ext ?_⟩ - rintro ⟨ab, b, rfl⟩ - simp [Set.range_inr] + simp measurable_invFun := Measurable.subtype_mk measurable_inr /-- Products distribute over sums (on the right) as measurable spaces. -/ diff --git a/Mathlib/MeasureTheory/MeasurableSpace/Pi.lean b/Mathlib/MeasureTheory/MeasurableSpace/Pi.lean new file mode 100644 index 0000000000000..82e1418f467fc --- /dev/null +++ b/Mathlib/MeasureTheory/MeasurableSpace/Pi.lean @@ -0,0 +1,107 @@ +/- +Copyright (c) 2020 Floris van Doorn. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Floris van Doorn +-/ +import Mathlib.MeasureTheory.MeasurableSpace.Basic +import Mathlib.MeasureTheory.PiSystem + +/-! +# Bases of the indexed product σ-algebra + +In this file we prove several versions of the following lemma: + +given a finite indexed collection of measurable spaces `α i`, +if the σ-algebra on each `α i` is generated by `C i`, +then the sets `{x | ∀ i, x i ∈ s i}`, where `s i ∈ C i`, +generate the σ-algebra on the indexed product of `α i`s. +-/ + +noncomputable section + +open Function Set Filter MeasurableSpace Encodable +open scoped Classical + +variable {ι : Type*} {α : ι → Type*} + +/-! We start with some measurability properties -/ + +lemma MeasurableSpace.pi_eq_generateFrom_projections {mα : ∀ i, MeasurableSpace (α i)} : + pi = generateFrom {B | ∃ (i : ι) (A : Set (α i)), MeasurableSet A ∧ eval i ⁻¹' A = B} := by + simp only [pi, ← generateFrom_iUnion_measurableSet, iUnion_setOf, measurableSet_comap] + +/-- Boxes formed by π-systems form a π-system. -/ +theorem IsPiSystem.pi {C : ∀ i, Set (Set (α i))} (hC : ∀ i, IsPiSystem (C i)) : + IsPiSystem (pi univ '' pi univ C) := by + rintro _ ⟨s₁, hs₁, rfl⟩ _ ⟨s₂, hs₂, rfl⟩ hst + rw [← pi_inter_distrib] at hst ⊢; rw [univ_pi_nonempty_iff] at hst + exact mem_image_of_mem _ fun i _ => hC i _ (hs₁ i (mem_univ i)) _ (hs₂ i (mem_univ i)) (hst i) + +/-- Boxes form a π-system. -/ +theorem isPiSystem_pi [∀ i, MeasurableSpace (α i)] : + IsPiSystem (pi univ '' pi univ fun i => { s : Set (α i) | MeasurableSet s }) := + IsPiSystem.pi fun _ => isPiSystem_measurableSet + +section Finite + +variable [Finite ι] + +/-- Boxes of countably spanning sets are countably spanning. -/ +theorem IsCountablySpanning.pi {C : ∀ i, Set (Set (α i))} (hC : ∀ i, IsCountablySpanning (C i)) : + IsCountablySpanning (pi univ '' pi univ C) := by + choose s h1s h2s using hC + cases nonempty_encodable (ι → ℕ) + let e : ℕ → ι → ℕ := fun n => (@decode (ι → ℕ) _ n).iget + refine ⟨fun n => Set.pi univ fun i => s i (e n i), fun n => + mem_image_of_mem _ fun i _ => h1s i _, ?_⟩ + simp_rw + [e, (surjective_decode_iget (ι → ℕ)).iUnion_comp fun x => Set.pi univ fun i => s i (x i), + iUnion_univ_pi s, h2s, pi_univ] + +/-- The product of generated σ-algebras is the one generated by boxes, if both generating sets + are countably spanning. -/ +theorem generateFrom_pi_eq {C : ∀ i, Set (Set (α i))} (hC : ∀ i, IsCountablySpanning (C i)) : + (@MeasurableSpace.pi _ _ fun i => generateFrom (C i)) = + generateFrom (pi univ '' pi univ C) := by + cases nonempty_encodable ι + apply le_antisymm + · refine iSup_le ?_; intro i; rw [comap_generateFrom] + apply generateFrom_le; rintro _ ⟨s, hs, rfl⟩ + choose t h1t h2t using hC + simp_rw [eval_preimage, ← h2t] + rw [← @iUnion_const _ ℕ _ s] + have : Set.pi univ (update (fun i' : ι => iUnion (t i')) i (⋃ _ : ℕ, s)) = + Set.pi univ fun k => ⋃ j : ℕ, + @update ι (fun i' => Set (α i')) _ (fun i' => t i' j) i s k := by + ext; simp_rw [mem_univ_pi]; apply forall_congr'; intro i' + by_cases h : i' = i + · subst h; simp + · rw [← Ne] at h; simp [h] + rw [this, ← iUnion_univ_pi] + apply MeasurableSet.iUnion + intro n; apply measurableSet_generateFrom + apply mem_image_of_mem; intro j _; dsimp only + by_cases h : j = i + · subst h; rwa [update_self] + · rw [update_of_ne h]; apply h1t + · apply generateFrom_le; rintro _ ⟨s, hs, rfl⟩ + rw [univ_pi_eq_iInter]; apply MeasurableSet.iInter; intro i + apply @measurable_pi_apply _ _ (fun i => generateFrom (C i)) + exact measurableSet_generateFrom (hs i (mem_univ i)) + +/-- If `C` and `D` generate the σ-algebras on `α` resp. `β`, then rectangles formed by `C` and `D` + generate the σ-algebra on `α × β`. -/ +theorem generateFrom_eq_pi [h : ∀ i, MeasurableSpace (α i)] {C : ∀ i, Set (Set (α i))} + (hC : ∀ i, generateFrom (C i) = h i) (h2C : ∀ i, IsCountablySpanning (C i)) : + generateFrom (pi univ '' pi univ C) = MeasurableSpace.pi := by + simp only [← funext hC, generateFrom_pi_eq h2C] + +/-- The product σ-algebra is generated from boxes, i.e. `s ×ˢ t` for sets `s : set α` and + `t : set β`. -/ +theorem generateFrom_pi [∀ i, MeasurableSpace (α i)] : + generateFrom (pi univ '' pi univ fun i => { s : Set (α i) | MeasurableSet s }) = + MeasurableSpace.pi := + generateFrom_eq_pi (fun _ => generateFrom_measurableSet) fun _ => + isCountablySpanning_measurableSet + +end Finite diff --git a/Mathlib/MeasureTheory/Measure/AbsolutelyContinuous.lean b/Mathlib/MeasureTheory/Measure/AbsolutelyContinuous.lean new file mode 100644 index 0000000000000..e201a358f0c78 --- /dev/null +++ b/Mathlib/MeasureTheory/Measure/AbsolutelyContinuous.lean @@ -0,0 +1,193 @@ +/- +Copyright (c) 2017 Johannes Hölzl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johannes Hölzl, Mario Carneiro +-/ +import Mathlib.MeasureTheory.Measure.Map + +/-! +# Absolute Continuity of Measures + +We say that `μ` is absolutely continuous with respect to `ν`, or that `μ` is dominated by `ν`, +if `ν(A) = 0` implies that `μ(A) = 0`. We denote that by `μ ≪ ν`. + +It is equivalent to an inequality of the almost everywhere filters of the measures: +`μ ≪ ν ↔ ae μ ≤ ae ν`. + +## Main definitions + +* `MeasureTheory.Measure.AbsolutelyContinuous μ ν`: `μ` is absolutely continuous with respect to `ν` + +## Main statements + +* `ae_le_iff_absolutelyContinuous`: `ae μ ≤ ae ν ↔ μ ≪ ν` + +## Notation + +* `μ ≪ ν`: `MeasureTheory.Measure.AbsolutelyContinuous μ ν`. That is: `μ` is absolutely continuous + with respect to `ν` + +-/ + +variable {α β δ ι R : Type*} + +namespace MeasureTheory + +open Set ENNReal NNReal + +variable {mα : MeasurableSpace α} {mβ : MeasurableSpace β} + {μ μ₁ μ₂ μ₃ ν ν' : Measure α} {s t : Set α} + +namespace Measure + +/-- We say that `μ` is absolutely continuous with respect to `ν`, or that `μ` is dominated by `ν`, + if `ν(A) = 0` implies that `μ(A) = 0`. -/ +def AbsolutelyContinuous {_m0 : MeasurableSpace α} (μ ν : Measure α) : Prop := + ∀ ⦃s : Set α⦄, ν s = 0 → μ s = 0 + +@[inherit_doc MeasureTheory.Measure.AbsolutelyContinuous] +scoped[MeasureTheory] infixl:50 " ≪ " => MeasureTheory.Measure.AbsolutelyContinuous + +theorem absolutelyContinuous_of_le (h : μ ≤ ν) : μ ≪ ν := fun s hs => + nonpos_iff_eq_zero.1 <| hs ▸ le_iff'.1 h s + +alias _root_.LE.le.absolutelyContinuous := absolutelyContinuous_of_le + +theorem absolutelyContinuous_of_eq (h : μ = ν) : μ ≪ ν := + h.le.absolutelyContinuous + +alias _root_.Eq.absolutelyContinuous := absolutelyContinuous_of_eq + +namespace AbsolutelyContinuous + +theorem mk (h : ∀ ⦃s : Set α⦄, MeasurableSet s → ν s = 0 → μ s = 0) : μ ≪ ν := by + intro s hs + rcases exists_measurable_superset_of_null hs with ⟨t, h1t, h2t, h3t⟩ + exact measure_mono_null h1t (h h2t h3t) + +@[refl] +protected theorem refl {_m0 : MeasurableSpace α} (μ : Measure α) : μ ≪ μ := + rfl.absolutelyContinuous + +protected theorem rfl : μ ≪ μ := fun _s hs => hs + +instance instIsRefl {_ : MeasurableSpace α} : IsRefl (Measure α) (· ≪ ·) := + ⟨fun _ => AbsolutelyContinuous.rfl⟩ + +@[simp] +protected lemma zero (μ : Measure α) : 0 ≪ μ := fun _ _ ↦ by simp + +@[trans] +protected theorem trans (h1 : μ₁ ≪ μ₂) (h2 : μ₂ ≪ μ₃) : μ₁ ≪ μ₃ := fun _s hs => h1 <| h2 hs + +@[mono] +protected theorem map (h : μ ≪ ν) {f : α → β} (hf : Measurable f) : μ.map f ≪ ν.map f := + AbsolutelyContinuous.mk fun s hs => by simpa [hf, hs] using @h _ + +protected theorem smul_left [SMul R ℝ≥0∞] [IsScalarTower R ℝ≥0∞ ℝ≥0∞] (h : μ ≪ ν) (c : R) : + c • μ ≪ ν := fun s hνs => by + simp only [h hνs, smul_apply, smul_zero, ← smul_one_smul ℝ≥0∞ c (0 : ℝ≥0∞)] + +/-- If `μ ≪ ν`, then `c • μ ≪ c • ν`. + +Earlier, this name was used for what's now called `AbsolutelyContinuous.smul_left`. -/ +protected theorem smul [SMul R ℝ≥0∞] [IsScalarTower R ℝ≥0∞ ℝ≥0∞] (h : μ ≪ ν) (c : R) : + c • μ ≪ c • ν := by + intro s hνs + rw [smul_apply, ← smul_one_smul ℝ≥0∞, smul_eq_mul, mul_eq_zero] at hνs ⊢ + exact hνs.imp_right fun hs ↦ h hs + +@[deprecated (since := "2024-11-14")] protected alias smul_both := AbsolutelyContinuous.smul + +protected lemma add (h1 : μ₁ ≪ ν) (h2 : μ₂ ≪ ν') : μ₁ + μ₂ ≪ ν + ν' := by + intro s hs + simp only [coe_add, Pi.add_apply, add_eq_zero] at hs ⊢ + exact ⟨h1 hs.1, h2 hs.2⟩ + +lemma add_left_iff {μ₁ μ₂ ν : Measure α} : + μ₁ + μ₂ ≪ ν ↔ μ₁ ≪ ν ∧ μ₂ ≪ ν := by + refine ⟨fun h ↦ ?_, fun h ↦ (h.1.add h.2).trans ?_⟩ + · have : ∀ s, ν s = 0 → μ₁ s = 0 ∧ μ₂ s = 0 := by intro s hs0; simpa using h hs0 + exact ⟨fun s hs0 ↦ (this s hs0).1, fun s hs0 ↦ (this s hs0).2⟩ + · rw [← two_smul ℝ≥0] + exact AbsolutelyContinuous.rfl.smul_left 2 + +lemma add_left {μ₁ μ₂ ν : Measure α} (h₁ : μ₁ ≪ ν) (h₂ : μ₂ ≪ ν) : μ₁ + μ₂ ≪ ν := + Measure.AbsolutelyContinuous.add_left_iff.mpr ⟨h₁, h₂⟩ + +lemma add_right (h1 : μ ≪ ν) (ν' : Measure α) : μ ≪ ν + ν' := by + intro s hs + simp only [coe_add, Pi.add_apply, add_eq_zero] at hs ⊢ + exact h1 hs.1 + +end AbsolutelyContinuous + +@[simp] +lemma absolutelyContinuous_zero_iff : μ ≪ 0 ↔ μ = 0 := + ⟨fun h ↦ measure_univ_eq_zero.mp (h rfl), fun h ↦ h.symm ▸ AbsolutelyContinuous.zero _⟩ + +alias absolutelyContinuous_refl := AbsolutelyContinuous.refl +alias absolutelyContinuous_rfl := AbsolutelyContinuous.rfl + +lemma absolutelyContinuous_sum_left {μs : ι → Measure α} (hμs : ∀ i, μs i ≪ ν) : + Measure.sum μs ≪ ν := + AbsolutelyContinuous.mk fun s hs hs0 ↦ by simp [sum_apply _ hs, fun i ↦ hμs i hs0] + +lemma absolutelyContinuous_sum_right {μs : ι → Measure α} (i : ι) (hνμ : ν ≪ μs i) : + ν ≪ Measure.sum μs := by + refine AbsolutelyContinuous.mk fun s hs hs0 ↦ ?_ + simp only [sum_apply _ hs, ENNReal.tsum_eq_zero] at hs0 + exact hνμ (hs0 i) + +lemma smul_absolutelyContinuous {c : ℝ≥0∞} : c • μ ≪ μ := .smul_left .rfl _ + +theorem absolutelyContinuous_of_le_smul {μ' : Measure α} {c : ℝ≥0∞} (hμ'_le : μ' ≤ c • μ) : + μ' ≪ μ := + (Measure.absolutelyContinuous_of_le hμ'_le).trans smul_absolutelyContinuous + +lemma absolutelyContinuous_smul {c : ℝ≥0∞} (hc : c ≠ 0) : μ ≪ c • μ := by + simp [AbsolutelyContinuous, hc] + +theorem ae_le_iff_absolutelyContinuous : ae μ ≤ ae ν ↔ μ ≪ ν := + ⟨fun h s => by + rw [measure_zero_iff_ae_nmem, measure_zero_iff_ae_nmem] + exact fun hs => h hs, fun h _ hs => h hs⟩ + +alias ⟨_root_.LE.le.absolutelyContinuous_of_ae, AbsolutelyContinuous.ae_le⟩ := + ae_le_iff_absolutelyContinuous + +alias ae_mono' := AbsolutelyContinuous.ae_le + +theorem AbsolutelyContinuous.ae_eq (h : μ ≪ ν) {f g : α → δ} (h' : f =ᵐ[ν] g) : f =ᵐ[μ] g := + h.ae_le h' + +end Measure + +protected theorem AEDisjoint.of_absolutelyContinuous + (h : AEDisjoint μ s t) {ν : Measure α} (h' : ν ≪ μ) : + AEDisjoint ν s t := h' h + +protected theorem AEDisjoint.of_le + (h : AEDisjoint μ s t) {ν : Measure α} (h' : ν ≤ μ) : + AEDisjoint ν s t := + h.of_absolutelyContinuous (Measure.absolutelyContinuous_of_le h') + +@[mono] +theorem ae_mono (h : μ ≤ ν) : ae μ ≤ ae ν := + h.absolutelyContinuous.ae_le + +end MeasureTheory + +namespace MeasurableEmbedding + +open MeasureTheory Measure + +variable {m0 : MeasurableSpace α} {m1 : MeasurableSpace β} {f : α → β} {μ ν : Measure α} + +lemma absolutelyContinuous_map (hf : MeasurableEmbedding f) (hμν : μ ≪ ν) : + μ.map f ≪ ν.map f := by + intro t ht + rw [hf.map_apply] at ht ⊢ + exact hμν ht + +end MeasurableEmbedding diff --git a/Mathlib/MeasureTheory/Measure/Comap.lean b/Mathlib/MeasureTheory/Measure/Comap.lean index 4b7a871e2f74c..457a53c721731 100644 --- a/Mathlib/MeasureTheory/Measure/Comap.lean +++ b/Mathlib/MeasureTheory/Measure/Comap.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov, Rémy Degenne -/ -import Mathlib.MeasureTheory.Measure.MeasureSpace +import Mathlib.MeasureTheory.Measure.Map /-! # Pullback of a measure diff --git a/Mathlib/MeasureTheory/Measure/Haar/Unique.lean b/Mathlib/MeasureTheory/Measure/Haar/Unique.lean index e9d2e87202f9e..a4ed925b7ac08 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/Unique.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/Unique.lean @@ -390,12 +390,6 @@ lemma haarScalarFactor_pos_of_isHaarMeasure (μ' μ : Measure G) [IsHaarMeasure [IsHaarMeasure μ'] : 0 < haarScalarFactor μ' μ := pos_iff_ne_zero.2 (fun H ↦ by simpa [H] using haarScalarFactor_eq_mul μ' μ μ') -@[deprecated (since := "2024-02-12")] -alias haarScalarFactor_pos_of_isOpenPosMeasure := haarScalarFactor_pos_of_isHaarMeasure - -@[deprecated (since := "2024-02-12")] -alias addHaarScalarFactor_pos_of_isOpenPosMeasure := addHaarScalarFactor_pos_of_isAddHaarMeasure - /-! ### Uniqueness of measure of sets with compact closure @@ -681,13 +675,6 @@ lemma isHaarMeasure_eq_of_isProbabilityMeasure [LocallyCompactSpace G] (μ' μ : ext s _hs simp [A s, ← Z] -@[deprecated (since := "2024-02-12")] -alias haarScalarFactor_eq_one_of_isProbabilityMeasure := isHaarMeasure_eq_of_isProbabilityMeasure - -@[deprecated (since := "2024-02-12")] -alias addHaarScalarFactor_eq_one_of_isProbabilityMeasure := - isAddHaarMeasure_eq_of_isProbabilityMeasure - /-! ### Uniqueness of measure of open sets @@ -896,8 +883,6 @@ lemma isMulLeftInvariant_eq_smul [LocallyCompactSpace G] [SecondCountableTopolog -- one could use as well `isMulLeftInvariant_eq_smul_of_innerRegular`, as in a -- second countable topological space all Haar measures are regular and inner regular -@[deprecated (since := "2024-02-12")] alias isHaarMeasure_eq_smul := isMulLeftInvariant_eq_smul -@[deprecated (since := "2024-02-12")] alias isAddHaarMeasure_eq_smul := isAddLeftInvariant_eq_smul /-- An invariant σ-finite measure is absolutely continuous with respect to a Haar measure in a second countable group. -/ diff --git a/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean b/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean index 2a68df2ddc985..aa7a1f5093442 100644 --- a/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean +++ b/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean @@ -128,6 +128,7 @@ namespace Measure ### Strict subspaces have zero measure -/ +open scoped Function -- required for scoped `on` notation /-- If a set is disjoint of its translates by infinitely many bounded vectors, then it has measure zero. This auxiliary lemma proves this assuming additionally that the set is bounded. -/ diff --git a/Mathlib/MeasureTheory/Measure/Map.lean b/Mathlib/MeasureTheory/Measure/Map.lean new file mode 100644 index 0000000000000..468fd36fa10a3 --- /dev/null +++ b/Mathlib/MeasureTheory/Measure/Map.lean @@ -0,0 +1,322 @@ +/- +Copyright (c) 2017 Johannes Hölzl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johannes Hölzl, Mario Carneiro +-/ +import Mathlib.MeasureTheory.MeasurableSpace.Embedding +import Mathlib.MeasureTheory.Measure.MeasureSpace + +/-! +# Pushforward of a measure + +In this file we define the pushforward `MeasureTheory.Measure.map f μ` +of a measure `μ` along an almost everywhere measurable map `f`. +If `f` is not a.e. measurable, then we define `map f μ` to be zero. + +## Main definitions + +* `MeasureTheory.Measure.map f μ`: map of the measure `μ` along the map `f`. + +## Main statements + +* `map_apply`: for `s` a measurable set, `μ.map f s = μ (f ⁻¹' s)` +* `map_map`: `(μ.map f).map g = μ.map (g ∘ f)` + +-/ + +variable {α β γ : Type*} + +open Set Function ENNReal NNReal +open Filter hiding map + +namespace MeasureTheory + +variable {mα : MeasurableSpace α} {mβ : MeasurableSpace β} {mγ : MeasurableSpace γ} + {μ ν : Measure α} {s : Set α} + +namespace Measure + +/-- Lift a linear map between `OuterMeasure` spaces such that for each measure `μ` every measurable +set is caratheodory-measurable w.r.t. `f μ` to a linear map between `Measure` spaces. -/ +noncomputable +def liftLinear [MeasurableSpace β] (f : OuterMeasure α →ₗ[ℝ≥0∞] OuterMeasure β) + (hf : ∀ μ : Measure α, ‹_› ≤ (f μ.toOuterMeasure).caratheodory) : + Measure α →ₗ[ℝ≥0∞] Measure β where + toFun μ := (f μ.toOuterMeasure).toMeasure (hf μ) + map_add' μ₁ μ₂ := ext fun s hs => by + simp only [map_add, coe_add, Pi.add_apply, toMeasure_apply, add_toOuterMeasure, + OuterMeasure.coe_add, hs] + map_smul' c μ := ext fun s hs => by + simp only [LinearMap.map_smulₛₗ, coe_smul, Pi.smul_apply, + toMeasure_apply, smul_toOuterMeasure (R := ℝ≥0∞), OuterMeasure.coe_smul (R := ℝ≥0∞), + smul_apply, hs] + +lemma liftLinear_apply₀ {f : OuterMeasure α →ₗ[ℝ≥0∞] OuterMeasure β} (hf) {s : Set β} + (hs : NullMeasurableSet s (liftLinear f hf μ)) : liftLinear f hf μ s = f μ.toOuterMeasure s := + toMeasure_apply₀ _ (hf μ) hs + +@[simp] +theorem liftLinear_apply {f : OuterMeasure α →ₗ[ℝ≥0∞] OuterMeasure β} (hf) {s : Set β} + (hs : MeasurableSet s) : liftLinear f hf μ s = f μ.toOuterMeasure s := + toMeasure_apply _ (hf μ) hs + +theorem le_liftLinear_apply {f : OuterMeasure α →ₗ[ℝ≥0∞] OuterMeasure β} (hf) (s : Set β) : + f μ.toOuterMeasure s ≤ liftLinear f hf μ s := + le_toMeasure_apply _ (hf μ) s + +open Classical in +/-- The pushforward of a measure as a linear map. It is defined to be `0` if `f` is not +a measurable function. -/ +noncomputable +def mapₗ [MeasurableSpace α] [MeasurableSpace β] (f : α → β) : Measure α →ₗ[ℝ≥0∞] Measure β := + if hf : Measurable f then + liftLinear (OuterMeasure.map f) fun μ _s hs t => + le_toOuterMeasure_caratheodory μ _ (hf hs) (f ⁻¹' t) + else 0 + +theorem mapₗ_congr {f g : α → β} (hf : Measurable f) (hg : Measurable g) (h : f =ᵐ[μ] g) : + mapₗ f μ = mapₗ g μ := by + ext1 s hs + simpa only [mapₗ, hf, hg, hs, dif_pos, liftLinear_apply, OuterMeasure.map_apply] + using measure_congr (h.preimage s) + +open Classical in +/-- The pushforward of a measure. It is defined to be `0` if `f` is not an almost everywhere +measurable function. -/ +noncomputable +irreducible_def map [MeasurableSpace α] [MeasurableSpace β] (f : α → β) (μ : Measure α) : + Measure β := + if hf : AEMeasurable f μ then mapₗ (hf.mk f) μ else 0 + +theorem mapₗ_mk_apply_of_aemeasurable {f : α → β} (hf : AEMeasurable f μ) : + mapₗ (hf.mk f) μ = map f μ := by simp [map, hf] + +theorem mapₗ_apply_of_measurable {f : α → β} (hf : Measurable f) (μ : Measure α) : + mapₗ f μ = map f μ := by + simp only [← mapₗ_mk_apply_of_aemeasurable hf.aemeasurable] + exact mapₗ_congr hf hf.aemeasurable.measurable_mk hf.aemeasurable.ae_eq_mk + +@[simp] +theorem map_add (μ ν : Measure α) {f : α → β} (hf : Measurable f) : + (μ + ν).map f = μ.map f + ν.map f := by simp [← mapₗ_apply_of_measurable hf] + +@[simp] +theorem map_zero (f : α → β) : (0 : Measure α).map f = 0 := by + by_cases hf : AEMeasurable f (0 : Measure α) <;> simp [map, hf] + +@[simp] +theorem map_of_not_aemeasurable {f : α → β} {μ : Measure α} (hf : ¬AEMeasurable f μ) : + μ.map f = 0 := by simp [map, hf] + +theorem _root_.AEMeasurable.of_map_ne_zero {f : α → β} {μ : Measure α} (hf : μ.map f ≠ 0) : + AEMeasurable f μ := not_imp_comm.1 map_of_not_aemeasurable hf + +theorem map_congr {f g : α → β} (h : f =ᵐ[μ] g) : Measure.map f μ = Measure.map g μ := by + by_cases hf : AEMeasurable f μ + · have hg : AEMeasurable g μ := hf.congr h + simp only [← mapₗ_mk_apply_of_aemeasurable hf, ← mapₗ_mk_apply_of_aemeasurable hg] + exact + mapₗ_congr hf.measurable_mk hg.measurable_mk (hf.ae_eq_mk.symm.trans (h.trans hg.ae_eq_mk)) + · have hg : ¬AEMeasurable g μ := by simpa [← aemeasurable_congr h] using hf + simp [map_of_not_aemeasurable, hf, hg] + +@[simp] +protected theorem map_smul {R : Type*} [SMul R ℝ≥0∞] [IsScalarTower R ℝ≥0∞ ℝ≥0∞] + (c : R) (μ : Measure α) (f : α → β) : (c • μ).map f = c • μ.map f := by + suffices ∀ c : ℝ≥0∞, (c • μ).map f = c • μ.map f by simpa using this (c • 1) + clear c; intro c + rcases eq_or_ne c 0 with (rfl | hc); · simp + by_cases hf : AEMeasurable f μ + · have hfc : AEMeasurable f (c • μ) := + ⟨hf.mk f, hf.measurable_mk, (ae_smul_measure_iff hc).2 hf.ae_eq_mk⟩ + simp only [← mapₗ_mk_apply_of_aemeasurable hf, ← mapₗ_mk_apply_of_aemeasurable hfc, + LinearMap.map_smulₛₗ, RingHom.id_apply] + congr 1 + apply mapₗ_congr hfc.measurable_mk hf.measurable_mk + exact EventuallyEq.trans ((ae_smul_measure_iff hc).1 hfc.ae_eq_mk.symm) hf.ae_eq_mk + · have hfc : ¬AEMeasurable f (c • μ) := by + intro hfc + exact hf ⟨hfc.mk f, hfc.measurable_mk, (ae_smul_measure_iff hc).1 hfc.ae_eq_mk⟩ + simp [map_of_not_aemeasurable hf, map_of_not_aemeasurable hfc] + + +@[deprecated Measure.map_smul (since := "2024-11-13")] +protected theorem map_smul_nnreal (c : ℝ≥0) (μ : Measure α) (f : α → β) : + (c • μ).map f = c • μ.map f := + μ.map_smul c f + +variable {f : α → β} + +lemma map_apply₀ {f : α → β} (hf : AEMeasurable f μ) {s : Set β} + (hs : NullMeasurableSet s (map f μ)) : μ.map f s = μ (f ⁻¹' s) := by + rw [map, dif_pos hf, mapₗ, dif_pos hf.measurable_mk] at hs ⊢ + rw [liftLinear_apply₀ _ hs, measure_congr (hf.ae_eq_mk.preimage s)] + rfl + +/-- We can evaluate the pushforward on measurable sets. For non-measurable sets, see + `MeasureTheory.Measure.le_map_apply` and `MeasurableEquiv.map_apply`. -/ +@[simp] +theorem map_apply_of_aemeasurable (hf : AEMeasurable f μ) {s : Set β} (hs : MeasurableSet s) : + μ.map f s = μ (f ⁻¹' s) := map_apply₀ hf hs.nullMeasurableSet + +@[simp] +theorem map_apply (hf : Measurable f) {s : Set β} (hs : MeasurableSet s) : + μ.map f s = μ (f ⁻¹' s) := + map_apply_of_aemeasurable hf.aemeasurable hs + +theorem map_toOuterMeasure (hf : AEMeasurable f μ) : + (μ.map f).toOuterMeasure = (OuterMeasure.map f μ.toOuterMeasure).trim := by + rw [← trimmed, OuterMeasure.trim_eq_trim_iff] + intro s hs + simp [hf, hs] + +@[simp] lemma map_eq_zero_iff (hf : AEMeasurable f μ) : μ.map f = 0 ↔ μ = 0 := by + simp_rw [← measure_univ_eq_zero, map_apply_of_aemeasurable hf .univ, preimage_univ] + +@[simp] lemma mapₗ_eq_zero_iff (hf : Measurable f) : Measure.mapₗ f μ = 0 ↔ μ = 0 := by + rw [mapₗ_apply_of_measurable hf, map_eq_zero_iff hf.aemeasurable] + +/-- If `map f μ = μ`, then the measure of the preimage of any null measurable set `s` +is equal to the measure of `s`. +Note that this lemma does not assume (a.e.) measurability of `f`. -/ +lemma measure_preimage_of_map_eq_self {f : α → α} (hf : map f μ = μ) + {s : Set α} (hs : NullMeasurableSet s μ) : μ (f ⁻¹' s) = μ s := by + if hfm : AEMeasurable f μ then + rw [← map_apply₀ hfm, hf] + rwa [hf] + else + rw [map_of_not_aemeasurable hfm] at hf + simp [← hf] + +lemma map_ne_zero_iff (hf : AEMeasurable f μ) : μ.map f ≠ 0 ↔ μ ≠ 0 := (map_eq_zero_iff hf).not +lemma mapₗ_ne_zero_iff (hf : Measurable f) : Measure.mapₗ f μ ≠ 0 ↔ μ ≠ 0 := + (mapₗ_eq_zero_iff hf).not + +@[simp] +theorem map_id : map id μ = μ := + ext fun _ => map_apply measurable_id + +@[simp] +theorem map_id' : map (fun x => x) μ = μ := + map_id + +/-- Mapping a measure twice is the same as mapping the measure with the composition. This version is +for measurable functions. See `map_map_of_aemeasurable` when they are just ae measurable. -/ +theorem map_map {g : β → γ} {f : α → β} (hg : Measurable g) (hf : Measurable f) : + (μ.map f).map g = μ.map (g ∘ f) := + ext fun s hs => by simp [hf, hg, hs, hg hs, hg.comp hf, ← preimage_comp] + +@[mono] +theorem map_mono {f : α → β} (h : μ ≤ ν) (hf : Measurable f) : μ.map f ≤ ν.map f := + le_iff.2 fun s hs ↦ by simp [hf.aemeasurable, hs, h _] + +/-- Even if `s` is not measurable, we can bound `map f μ s` from below. + See also `MeasurableEquiv.map_apply`. -/ +theorem le_map_apply {f : α → β} (hf : AEMeasurable f μ) (s : Set β) : μ (f ⁻¹' s) ≤ μ.map f s := + calc + μ (f ⁻¹' s) ≤ μ (f ⁻¹' toMeasurable (μ.map f) s) := by gcongr; apply subset_toMeasurable + _ = μ.map f (toMeasurable (μ.map f) s) := + (map_apply_of_aemeasurable hf <| measurableSet_toMeasurable _ _).symm + _ = μ.map f s := measure_toMeasurable _ + +theorem le_map_apply_image {f : α → β} (hf : AEMeasurable f μ) (s : Set α) : + μ s ≤ μ.map f (f '' s) := + (measure_mono (subset_preimage_image f s)).trans (le_map_apply hf _) + +/-- Even if `s` is not measurable, `map f μ s = 0` implies that `μ (f ⁻¹' s) = 0`. -/ +theorem preimage_null_of_map_null {f : α → β} (hf : AEMeasurable f μ) {s : Set β} + (hs : μ.map f s = 0) : μ (f ⁻¹' s) = 0 := + nonpos_iff_eq_zero.mp <| (le_map_apply hf s).trans_eq hs + +theorem tendsto_ae_map {f : α → β} (hf : AEMeasurable f μ) : Tendsto f (ae μ) (ae (μ.map f)) := + fun _ hs => preimage_null_of_map_null hf hs + +end Measure + +open Measure + +theorem mem_ae_map_iff {f : α → β} (hf : AEMeasurable f μ) {s : Set β} (hs : MeasurableSet s) : + s ∈ ae (μ.map f) ↔ f ⁻¹' s ∈ ae μ := by + simp only [mem_ae_iff, map_apply_of_aemeasurable hf hs.compl, preimage_compl] + +theorem mem_ae_of_mem_ae_map {f : α → β} (hf : AEMeasurable f μ) {s : Set β} + (hs : s ∈ ae (μ.map f)) : f ⁻¹' s ∈ ae μ := + (tendsto_ae_map hf).eventually hs + +theorem ae_map_iff {f : α → β} (hf : AEMeasurable f μ) {p : β → Prop} + (hp : MeasurableSet { x | p x }) : (∀ᵐ y ∂μ.map f, p y) ↔ ∀ᵐ x ∂μ, p (f x) := + mem_ae_map_iff hf hp + +theorem ae_of_ae_map {f : α → β} (hf : AEMeasurable f μ) {p : β → Prop} (h : ∀ᵐ y ∂μ.map f, p y) : + ∀ᵐ x ∂μ, p (f x) := + mem_ae_of_mem_ae_map hf h + +theorem ae_map_mem_range {m0 : MeasurableSpace α} (f : α → β) (hf : MeasurableSet (range f)) + (μ : Measure α) : ∀ᵐ x ∂μ.map f, x ∈ range f := by + by_cases h : AEMeasurable f μ + · change range f ∈ ae (μ.map f) + rw [mem_ae_map_iff h hf] + filter_upwards using mem_range_self + · simp [map_of_not_aemeasurable h] + +end MeasureTheory + +namespace MeasurableEmbedding + +open MeasureTheory Measure + +variable {m0 : MeasurableSpace α} {m1 : MeasurableSpace β} {f : α → β} {μ ν : Measure α} + +nonrec theorem map_apply (hf : MeasurableEmbedding f) (μ : Measure α) (s : Set β) : + μ.map f s = μ (f ⁻¹' s) := by + refine le_antisymm ?_ (le_map_apply hf.measurable.aemeasurable s) + set t := f '' toMeasurable μ (f ⁻¹' s) ∪ (range f)ᶜ + have htm : MeasurableSet t := + (hf.measurableSet_image.2 <| measurableSet_toMeasurable _ _).union + hf.measurableSet_range.compl + have hst : s ⊆ t := by + rw [subset_union_compl_iff_inter_subset, ← image_preimage_eq_inter_range] + exact image_subset _ (subset_toMeasurable _ _) + have hft : f ⁻¹' t = toMeasurable μ (f ⁻¹' s) := by + rw [preimage_union, preimage_compl, preimage_range, compl_univ, union_empty, + hf.injective.preimage_image] + calc + μ.map f s ≤ μ.map f t := by gcongr + _ = μ (f ⁻¹' s) := by rw [map_apply hf.measurable htm, hft, measure_toMeasurable] + +end MeasurableEmbedding + +namespace MeasurableEquiv + +/-! Interactions of measurable equivalences and measures -/ + +open Equiv MeasureTheory MeasureTheory.Measure + +variable {_ : MeasurableSpace α} [MeasurableSpace β] {μ : Measure α} {ν : Measure β} + +/-- If we map a measure along a measurable equivalence, we can compute the measure on all sets + (not just the measurable ones). -/ +protected theorem map_apply (f : α ≃ᵐ β) (s : Set β) : μ.map f s = μ (f ⁻¹' s) := + f.measurableEmbedding.map_apply _ _ + +@[simp] +theorem map_symm_map (e : α ≃ᵐ β) : (μ.map e).map e.symm = μ := by + simp [map_map e.symm.measurable e.measurable] + +@[simp] +theorem map_map_symm (e : α ≃ᵐ β) : (ν.map e.symm).map e = ν := by + simp [map_map e.measurable e.symm.measurable] + +theorem map_measurableEquiv_injective (e : α ≃ᵐ β) : Injective (Measure.map e) := by + intro μ₁ μ₂ hμ + apply_fun Measure.map e.symm at hμ + simpa [map_symm_map e] using hμ + +theorem map_apply_eq_iff_map_symm_apply_eq (e : α ≃ᵐ β) : μ.map e = ν ↔ ν.map e.symm = μ := by + rw [← (map_measurableEquiv_injective e).eq_iff, map_map_symm, eq_comm] + +theorem map_ae (f : α ≃ᵐ β) (μ : Measure α) : Filter.map f (ae μ) = ae (map f μ) := by + ext s + simp_rw [mem_map, mem_ae_iff, ← preimage_compl, f.map_apply] + +end MeasurableEquiv diff --git a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean index a389edbf53a51..54c1d8c77d045 100644 --- a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean +++ b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean @@ -4,8 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro -/ import Mathlib.MeasureTheory.Measure.NullMeasurable -import Mathlib.MeasureTheory.MeasurableSpace.Embedding -import Mathlib.MeasureTheory.OuterMeasure.BorelCantelli import Mathlib.Topology.Algebra.Order.LiminfLimsup import Mathlib.Order.Interval.Set.Monotone @@ -626,6 +624,24 @@ theorem tendsto_measure_iInter_le {α ι : Type*} {_ : MeasurableSpace α} {μ : exact tendsto_atTop_iInf fun i j hij ↦ measure_mono <| biInter_subset_biInter_left fun k hki ↦ le_trans hki hij +/-- Some version of continuity of a measure in the emptyset using the intersection along a set of +sets. -/ +theorem exists_measure_iInter_lt {α ι : Type*} {_ : MeasurableSpace α} {μ : Measure α} + [SemilatticeSup ι] [Countable ι] {f : ι → Set α} + (hm : ∀ i, NullMeasurableSet (f i) μ) {ε : ℝ≥0∞} (hε : 0 < ε) (hfin : ∃ i, μ (f i) ≠ ∞) + (hfem : ⋂ n, f n = ∅) : ∃ m, μ (⋂ n ≤ m, f n) < ε := by + let F m := μ (⋂ n ≤ m, f n) + have hFAnti : Antitone F := + fun i j hij => measure_mono (biInter_subset_biInter_left fun k hki => le_trans hki hij) + suffices Filter.Tendsto F Filter.atTop (𝓝 0) by + rw [@ENNReal.tendsto_atTop_zero_iff_lt_of_antitone + _ (nonempty_of_exists hfin) _ _ hFAnti] at this + exact this ε hε + have hzero : μ (⋂ n, f n) = 0 := by + simp only [hfem, measure_empty] + rw [← hzero] + exact tendsto_measure_iInter_le hm hfin + /-- The measure of the intersection of a decreasing sequence of measurable sets indexed by a linear order with first countable topology is the limit of the measures. -/ theorem tendsto_measure_biInter_gt {ι : Type*} [LinearOrder ι] [TopologicalSpace ι] @@ -1122,201 +1138,6 @@ lemma nonempty_of_neZero (μ : Measure α) [NeZero μ] : Nonempty α := (isEmpty_or_nonempty α).resolve_left fun h ↦ by simpa [eq_empty_of_isEmpty] using NeZero.ne (μ univ) -/-! ### Pushforward and pullback -/ - - -/-- Lift a linear map between `OuterMeasure` spaces such that for each measure `μ` every measurable -set is caratheodory-measurable w.r.t. `f μ` to a linear map between `Measure` spaces. -/ -def liftLinear [MeasurableSpace β] (f : OuterMeasure α →ₗ[ℝ≥0∞] OuterMeasure β) - (hf : ∀ μ : Measure α, ‹_› ≤ (f μ.toOuterMeasure).caratheodory) : - Measure α →ₗ[ℝ≥0∞] Measure β where - toFun μ := (f μ.toOuterMeasure).toMeasure (hf μ) - map_add' μ₁ μ₂ := ext fun s hs => by - simp only [map_add, coe_add, Pi.add_apply, toMeasure_apply, add_toOuterMeasure, - OuterMeasure.coe_add, hs] - map_smul' c μ := ext fun s hs => by - simp only [LinearMap.map_smulₛₗ, coe_smul, Pi.smul_apply, - toMeasure_apply, smul_toOuterMeasure (R := ℝ≥0∞), OuterMeasure.coe_smul (R := ℝ≥0∞), - smul_apply, hs] - -lemma liftLinear_apply₀ {f : OuterMeasure α →ₗ[ℝ≥0∞] OuterMeasure β} (hf) {s : Set β} - (hs : NullMeasurableSet s (liftLinear f hf μ)) : liftLinear f hf μ s = f μ.toOuterMeasure s := - toMeasure_apply₀ _ (hf μ) hs - -@[simp] -theorem liftLinear_apply {f : OuterMeasure α →ₗ[ℝ≥0∞] OuterMeasure β} (hf) {s : Set β} - (hs : MeasurableSet s) : liftLinear f hf μ s = f μ.toOuterMeasure s := - toMeasure_apply _ (hf μ) hs - -theorem le_liftLinear_apply {f : OuterMeasure α →ₗ[ℝ≥0∞] OuterMeasure β} (hf) (s : Set β) : - f μ.toOuterMeasure s ≤ liftLinear f hf μ s := - le_toMeasure_apply _ (hf μ) s - -open Classical in -/-- The pushforward of a measure as a linear map. It is defined to be `0` if `f` is not -a measurable function. -/ -def mapₗ [MeasurableSpace α] [MeasurableSpace β] (f : α → β) : Measure α →ₗ[ℝ≥0∞] Measure β := - if hf : Measurable f then - liftLinear (OuterMeasure.map f) fun μ _s hs t => - le_toOuterMeasure_caratheodory μ _ (hf hs) (f ⁻¹' t) - else 0 - -theorem mapₗ_congr {f g : α → β} (hf : Measurable f) (hg : Measurable g) (h : f =ᵐ[μ] g) : - mapₗ f μ = mapₗ g μ := by - ext1 s hs - simpa only [mapₗ, hf, hg, hs, dif_pos, liftLinear_apply, OuterMeasure.map_apply] - using measure_congr (h.preimage s) - -open Classical in -/-- The pushforward of a measure. It is defined to be `0` if `f` is not an almost everywhere -measurable function. -/ -irreducible_def map [MeasurableSpace α] [MeasurableSpace β] (f : α → β) (μ : Measure α) : - Measure β := - if hf : AEMeasurable f μ then mapₗ (hf.mk f) μ else 0 - -theorem mapₗ_mk_apply_of_aemeasurable {f : α → β} (hf : AEMeasurable f μ) : - mapₗ (hf.mk f) μ = map f μ := by simp [map, hf] - -theorem mapₗ_apply_of_measurable {f : α → β} (hf : Measurable f) (μ : Measure α) : - mapₗ f μ = map f μ := by - simp only [← mapₗ_mk_apply_of_aemeasurable hf.aemeasurable] - exact mapₗ_congr hf hf.aemeasurable.measurable_mk hf.aemeasurable.ae_eq_mk - -@[simp] -theorem map_add (μ ν : Measure α) {f : α → β} (hf : Measurable f) : - (μ + ν).map f = μ.map f + ν.map f := by simp [← mapₗ_apply_of_measurable hf] - -@[simp] -theorem map_zero (f : α → β) : (0 : Measure α).map f = 0 := by - by_cases hf : AEMeasurable f (0 : Measure α) <;> simp [map, hf] - -@[simp] -theorem map_of_not_aemeasurable {f : α → β} {μ : Measure α} (hf : ¬AEMeasurable f μ) : - μ.map f = 0 := by simp [map, hf] - -theorem _root_.AEMeasurable.of_map_ne_zero {f : α → β} {μ : Measure α} (hf : μ.map f ≠ 0) : - AEMeasurable f μ := not_imp_comm.1 map_of_not_aemeasurable hf - -theorem map_congr {f g : α → β} (h : f =ᵐ[μ] g) : Measure.map f μ = Measure.map g μ := by - by_cases hf : AEMeasurable f μ - · have hg : AEMeasurable g μ := hf.congr h - simp only [← mapₗ_mk_apply_of_aemeasurable hf, ← mapₗ_mk_apply_of_aemeasurable hg] - exact - mapₗ_congr hf.measurable_mk hg.measurable_mk (hf.ae_eq_mk.symm.trans (h.trans hg.ae_eq_mk)) - · have hg : ¬AEMeasurable g μ := by simpa [← aemeasurable_congr h] using hf - simp [map_of_not_aemeasurable, hf, hg] - -@[simp] -protected theorem map_smul {R : Type*} [SMul R ℝ≥0∞] [IsScalarTower R ℝ≥0∞ ℝ≥0∞] - (c : R) (μ : Measure α) (f : α → β) : (c • μ).map f = c • μ.map f := by - suffices ∀ c : ℝ≥0∞, (c • μ).map f = c • μ.map f by simpa using this (c • 1) - clear c; intro c - rcases eq_or_ne c 0 with (rfl | hc); · simp - by_cases hf : AEMeasurable f μ - · have hfc : AEMeasurable f (c • μ) := - ⟨hf.mk f, hf.measurable_mk, (ae_smul_measure_iff hc).2 hf.ae_eq_mk⟩ - simp only [← mapₗ_mk_apply_of_aemeasurable hf, ← mapₗ_mk_apply_of_aemeasurable hfc, - LinearMap.map_smulₛₗ, RingHom.id_apply] - congr 1 - apply mapₗ_congr hfc.measurable_mk hf.measurable_mk - exact EventuallyEq.trans ((ae_smul_measure_iff hc).1 hfc.ae_eq_mk.symm) hf.ae_eq_mk - · have hfc : ¬AEMeasurable f (c • μ) := by - intro hfc - exact hf ⟨hfc.mk f, hfc.measurable_mk, (ae_smul_measure_iff hc).1 hfc.ae_eq_mk⟩ - simp [map_of_not_aemeasurable hf, map_of_not_aemeasurable hfc] - - -@[deprecated Measure.map_smul (since := "2024-11-13")] -protected theorem map_smul_nnreal (c : ℝ≥0) (μ : Measure α) (f : α → β) : - (c • μ).map f = c • μ.map f := - μ.map_smul c f - -variable {f : α → β} - -lemma map_apply₀ {f : α → β} (hf : AEMeasurable f μ) {s : Set β} - (hs : NullMeasurableSet s (map f μ)) : μ.map f s = μ (f ⁻¹' s) := by - rw [map, dif_pos hf, mapₗ, dif_pos hf.measurable_mk] at hs ⊢ - rw [liftLinear_apply₀ _ hs, measure_congr (hf.ae_eq_mk.preimage s)] - rfl - -/-- We can evaluate the pushforward on measurable sets. For non-measurable sets, see - `MeasureTheory.Measure.le_map_apply` and `MeasurableEquiv.map_apply`. -/ -@[simp] -theorem map_apply_of_aemeasurable (hf : AEMeasurable f μ) {s : Set β} (hs : MeasurableSet s) : - μ.map f s = μ (f ⁻¹' s) := map_apply₀ hf hs.nullMeasurableSet - -@[simp] -theorem map_apply (hf : Measurable f) {s : Set β} (hs : MeasurableSet s) : - μ.map f s = μ (f ⁻¹' s) := - map_apply_of_aemeasurable hf.aemeasurable hs - -theorem map_toOuterMeasure (hf : AEMeasurable f μ) : - (μ.map f).toOuterMeasure = (OuterMeasure.map f μ.toOuterMeasure).trim := by - rw [← trimmed, OuterMeasure.trim_eq_trim_iff] - intro s hs - simp [hf, hs] - -@[simp] lemma map_eq_zero_iff (hf : AEMeasurable f μ) : μ.map f = 0 ↔ μ = 0 := by - simp_rw [← measure_univ_eq_zero, map_apply_of_aemeasurable hf .univ, preimage_univ] - -@[simp] lemma mapₗ_eq_zero_iff (hf : Measurable f) : Measure.mapₗ f μ = 0 ↔ μ = 0 := by - rw [mapₗ_apply_of_measurable hf, map_eq_zero_iff hf.aemeasurable] - -/-- If `map f μ = μ`, then the measure of the preimage of any null measurable set `s` -is equal to the measure of `s`. -Note that this lemma does not assume (a.e.) measurability of `f`. -/ -lemma measure_preimage_of_map_eq_self {f : α → α} (hf : map f μ = μ) - {s : Set α} (hs : NullMeasurableSet s μ) : μ (f ⁻¹' s) = μ s := by - if hfm : AEMeasurable f μ then - rw [← map_apply₀ hfm, hf] - rwa [hf] - else - rw [map_of_not_aemeasurable hfm] at hf - simp [← hf] - -lemma map_ne_zero_iff (hf : AEMeasurable f μ) : μ.map f ≠ 0 ↔ μ ≠ 0 := (map_eq_zero_iff hf).not -lemma mapₗ_ne_zero_iff (hf : Measurable f) : Measure.mapₗ f μ ≠ 0 ↔ μ ≠ 0 := - (mapₗ_eq_zero_iff hf).not - -@[simp] -theorem map_id : map id μ = μ := - ext fun _ => map_apply measurable_id - -@[simp] -theorem map_id' : map (fun x => x) μ = μ := - map_id - -/-- Mapping a measure twice is the same as mapping the measure with the composition. This version is -for measurable functions. See `map_map_of_aemeasurable` when they are just ae measurable. -/ -theorem map_map {g : β → γ} {f : α → β} (hg : Measurable g) (hf : Measurable f) : - (μ.map f).map g = μ.map (g ∘ f) := - ext fun s hs => by simp [hf, hg, hs, hg hs, hg.comp hf, ← preimage_comp] - -@[mono] -theorem map_mono {f : α → β} (h : μ ≤ ν) (hf : Measurable f) : μ.map f ≤ ν.map f := - le_iff.2 fun s hs ↦ by simp [hf.aemeasurable, hs, h _] - -/-- Even if `s` is not measurable, we can bound `map f μ s` from below. - See also `MeasurableEquiv.map_apply`. -/ -theorem le_map_apply {f : α → β} (hf : AEMeasurable f μ) (s : Set β) : μ (f ⁻¹' s) ≤ μ.map f s := - calc - μ (f ⁻¹' s) ≤ μ (f ⁻¹' toMeasurable (μ.map f) s) := by gcongr; apply subset_toMeasurable - _ = μ.map f (toMeasurable (μ.map f) s) := - (map_apply_of_aemeasurable hf <| measurableSet_toMeasurable _ _).symm - _ = μ.map f s := measure_toMeasurable _ - -theorem le_map_apply_image {f : α → β} (hf : AEMeasurable f μ) (s : Set α) : - μ s ≤ μ.map f (f '' s) := - (measure_mono (subset_preimage_image f s)).trans (le_map_apply hf _) - -/-- Even if `s` is not measurable, `map f μ s = 0` implies that `μ (f ⁻¹' s) = 0`. -/ -theorem preimage_null_of_map_null {f : α → β} (hf : AEMeasurable f μ) {s : Set β} - (hs : μ.map f s = 0) : μ (f ⁻¹' s) = 0 := - nonpos_iff_eq_zero.mp <| (le_map_apply hf s).trans_eq hs - -theorem tendsto_ae_map {f : α → β} (hf : AEMeasurable f μ) : Tendsto f (ae μ) (ae (μ.map f)) := - fun _ hs => preimage_null_of_map_null hf hs - section Sum variable {f : ι → Measure α} @@ -1448,307 +1269,6 @@ theorem sum_add_sum {ι : Type*} (μ ν : ι → Measure α) : sum μ + sum ν = simp [*, Function.apply_extend (fun μ : Measure α ↦ μ s)] end Sum -/-! ### Absolute continuity -/ - -/-- We say that `μ` is absolutely continuous with respect to `ν`, or that `μ` is dominated by `ν`, - if `ν(A) = 0` implies that `μ(A) = 0`. -/ -def AbsolutelyContinuous {_m0 : MeasurableSpace α} (μ ν : Measure α) : Prop := - ∀ ⦃s : Set α⦄, ν s = 0 → μ s = 0 - -@[inherit_doc MeasureTheory.Measure.AbsolutelyContinuous] -scoped[MeasureTheory] infixl:50 " ≪ " => MeasureTheory.Measure.AbsolutelyContinuous - -theorem absolutelyContinuous_of_le (h : μ ≤ ν) : μ ≪ ν := fun s hs => - nonpos_iff_eq_zero.1 <| hs ▸ le_iff'.1 h s - -alias _root_.LE.le.absolutelyContinuous := absolutelyContinuous_of_le - -theorem absolutelyContinuous_of_eq (h : μ = ν) : μ ≪ ν := - h.le.absolutelyContinuous - -alias _root_.Eq.absolutelyContinuous := absolutelyContinuous_of_eq - -namespace AbsolutelyContinuous - -theorem mk (h : ∀ ⦃s : Set α⦄, MeasurableSet s → ν s = 0 → μ s = 0) : μ ≪ ν := by - intro s hs - rcases exists_measurable_superset_of_null hs with ⟨t, h1t, h2t, h3t⟩ - exact measure_mono_null h1t (h h2t h3t) - -@[refl] -protected theorem refl {_m0 : MeasurableSpace α} (μ : Measure α) : μ ≪ μ := - rfl.absolutelyContinuous - -protected theorem rfl : μ ≪ μ := fun _s hs => hs - -instance instIsRefl {_ : MeasurableSpace α} : IsRefl (Measure α) (· ≪ ·) := - ⟨fun _ => AbsolutelyContinuous.rfl⟩ - -@[simp] -protected lemma zero (μ : Measure α) : 0 ≪ μ := fun _ _ ↦ by simp - -@[trans] -protected theorem trans (h1 : μ₁ ≪ μ₂) (h2 : μ₂ ≪ μ₃) : μ₁ ≪ μ₃ := fun _s hs => h1 <| h2 hs - -@[mono] -protected theorem map (h : μ ≪ ν) {f : α → β} (hf : Measurable f) : μ.map f ≪ ν.map f := - AbsolutelyContinuous.mk fun s hs => by simpa [hf, hs] using @h _ - -protected theorem smul_left [SMul R ℝ≥0∞] [IsScalarTower R ℝ≥0∞ ℝ≥0∞] (h : μ ≪ ν) (c : R) : - c • μ ≪ ν := fun s hνs => by - simp only [h hνs, smul_apply, smul_zero, ← smul_one_smul ℝ≥0∞ c (0 : ℝ≥0∞)] - -/-- If `μ ≪ ν`, then `c • μ ≪ c • ν`. - -Earlier, this name was used for what's now called `AbsolutelyContinuous.smul_left`. -/ -protected theorem smul [SMul R ℝ≥0∞] [IsScalarTower R ℝ≥0∞ ℝ≥0∞] (h : μ ≪ ν) (c : R) : - c • μ ≪ c • ν := by - intro s hνs - rw [smul_apply, ← smul_one_smul ℝ≥0∞, smul_eq_mul, mul_eq_zero] at hνs ⊢ - exact hνs.imp_right fun hs ↦ h hs - -@[deprecated (since := "2024-11-14")] protected alias smul_both := AbsolutelyContinuous.smul - -protected lemma add (h1 : μ₁ ≪ ν) (h2 : μ₂ ≪ ν') : μ₁ + μ₂ ≪ ν + ν' := by - intro s hs - simp only [coe_add, Pi.add_apply, add_eq_zero] at hs ⊢ - exact ⟨h1 hs.1, h2 hs.2⟩ - -lemma add_left_iff {μ₁ μ₂ ν : Measure α} : - μ₁ + μ₂ ≪ ν ↔ μ₁ ≪ ν ∧ μ₂ ≪ ν := by - refine ⟨fun h ↦ ?_, fun h ↦ (h.1.add h.2).trans ?_⟩ - · have : ∀ s, ν s = 0 → μ₁ s = 0 ∧ μ₂ s = 0 := by intro s hs0; simpa using h hs0 - exact ⟨fun s hs0 ↦ (this s hs0).1, fun s hs0 ↦ (this s hs0).2⟩ - · rw [← two_smul ℝ≥0] - exact AbsolutelyContinuous.rfl.smul_left 2 - -lemma add_left {μ₁ μ₂ ν : Measure α} (h₁ : μ₁ ≪ ν) (h₂ : μ₂ ≪ ν) : μ₁ + μ₂ ≪ ν := - Measure.AbsolutelyContinuous.add_left_iff.mpr ⟨h₁, h₂⟩ - -lemma add_right (h1 : μ ≪ ν) (ν' : Measure α) : μ ≪ ν + ν' := by - intro s hs - simp only [coe_add, Pi.add_apply, add_eq_zero] at hs ⊢ - exact h1 hs.1 - -end AbsolutelyContinuous - -@[simp] -lemma absolutelyContinuous_zero_iff : μ ≪ 0 ↔ μ = 0 := - ⟨fun h ↦ measure_univ_eq_zero.mp (h rfl), fun h ↦ h.symm ▸ AbsolutelyContinuous.zero _⟩ - -alias absolutelyContinuous_refl := AbsolutelyContinuous.refl -alias absolutelyContinuous_rfl := AbsolutelyContinuous.rfl - -lemma absolutelyContinuous_sum_left {μs : ι → Measure α} (hμs : ∀ i, μs i ≪ ν) : - Measure.sum μs ≪ ν := - AbsolutelyContinuous.mk fun s hs hs0 ↦ by simp [sum_apply _ hs, fun i ↦ hμs i hs0] - -lemma absolutelyContinuous_sum_right {μs : ι → Measure α} (i : ι) (hνμ : ν ≪ μs i) : - ν ≪ Measure.sum μs := by - refine AbsolutelyContinuous.mk fun s hs hs0 ↦ ?_ - simp only [sum_apply _ hs, ENNReal.tsum_eq_zero] at hs0 - exact hνμ (hs0 i) - -lemma smul_absolutelyContinuous {c : ℝ≥0∞} : c • μ ≪ μ := .smul_left .rfl _ - -theorem absolutelyContinuous_of_le_smul {μ' : Measure α} {c : ℝ≥0∞} (hμ'_le : μ' ≤ c • μ) : - μ' ≪ μ := - (Measure.absolutelyContinuous_of_le hμ'_le).trans smul_absolutelyContinuous - -lemma absolutelyContinuous_smul {c : ℝ≥0∞} (hc : c ≠ 0) : μ ≪ c • μ := by - simp [AbsolutelyContinuous, hc] - -theorem ae_le_iff_absolutelyContinuous : ae μ ≤ ae ν ↔ μ ≪ ν := - ⟨fun h s => by - rw [measure_zero_iff_ae_nmem, measure_zero_iff_ae_nmem] - exact fun hs => h hs, fun h _ hs => h hs⟩ - -alias ⟨_root_.LE.le.absolutelyContinuous_of_ae, AbsolutelyContinuous.ae_le⟩ := - ae_le_iff_absolutelyContinuous - -alias ae_mono' := AbsolutelyContinuous.ae_le - -theorem AbsolutelyContinuous.ae_eq (h : μ ≪ ν) {f g : α → δ} (h' : f =ᵐ[ν] g) : f =ᵐ[μ] g := - h.ae_le h' - -protected theorem _root_.MeasureTheory.AEDisjoint.of_absolutelyContinuous - (h : AEDisjoint μ s t) {ν : Measure α} (h' : ν ≪ μ) : - AEDisjoint ν s t := h' h - -protected theorem _root_.MeasureTheory.AEDisjoint.of_le - (h : AEDisjoint μ s t) {ν : Measure α} (h' : ν ≤ μ) : - AEDisjoint ν s t := - h.of_absolutelyContinuous (Measure.absolutelyContinuous_of_le h') - -/-! ### Quasi measure preserving maps (a.k.a. non-singular maps) -/ - - -/-- A map `f : α → β` is said to be *quasi measure preserving* (a.k.a. non-singular) w.r.t. measures -`μa` and `μb` if it is measurable and `μb s = 0` implies `μa (f ⁻¹' s) = 0`. -/ -structure QuasiMeasurePreserving {m0 : MeasurableSpace α} (f : α → β) - (μa : Measure α := by volume_tac) - (μb : Measure β := by volume_tac) : Prop where - protected measurable : Measurable f - protected absolutelyContinuous : μa.map f ≪ μb - -namespace QuasiMeasurePreserving - -protected theorem id {_m0 : MeasurableSpace α} (μ : Measure α) : QuasiMeasurePreserving id μ μ := - ⟨measurable_id, map_id.absolutelyContinuous⟩ - -variable {μa μa' : Measure α} {μb μb' : Measure β} {μc : Measure γ} {f : α → β} - -protected theorem _root_.Measurable.quasiMeasurePreserving - {_m0 : MeasurableSpace α} (hf : Measurable f) (μ : Measure α) : - QuasiMeasurePreserving f μ (μ.map f) := - ⟨hf, AbsolutelyContinuous.rfl⟩ - -theorem mono_left (h : QuasiMeasurePreserving f μa μb) (ha : μa' ≪ μa) : - QuasiMeasurePreserving f μa' μb := - ⟨h.1, (ha.map h.1).trans h.2⟩ - -theorem mono_right (h : QuasiMeasurePreserving f μa μb) (ha : μb ≪ μb') : - QuasiMeasurePreserving f μa μb' := - ⟨h.1, h.2.trans ha⟩ - -@[mono] -theorem mono (ha : μa' ≪ μa) (hb : μb ≪ μb') (h : QuasiMeasurePreserving f μa μb) : - QuasiMeasurePreserving f μa' μb' := - (h.mono_left ha).mono_right hb - -protected theorem comp {g : β → γ} {f : α → β} (hg : QuasiMeasurePreserving g μb μc) - (hf : QuasiMeasurePreserving f μa μb) : QuasiMeasurePreserving (g ∘ f) μa μc := - ⟨hg.measurable.comp hf.measurable, by - rw [← map_map hg.1 hf.1] - exact (hf.2.map hg.1).trans hg.2⟩ - -protected theorem iterate {f : α → α} (hf : QuasiMeasurePreserving f μa μa) : - ∀ n, QuasiMeasurePreserving f^[n] μa μa - | 0 => QuasiMeasurePreserving.id μa - | n + 1 => (hf.iterate n).comp hf - -protected theorem aemeasurable (hf : QuasiMeasurePreserving f μa μb) : AEMeasurable f μa := - hf.1.aemeasurable - -theorem smul_measure {R : Type*} [SMul R ℝ≥0∞] [IsScalarTower R ℝ≥0∞ ℝ≥0∞] - (hf : QuasiMeasurePreserving f μa μb) (c : R) : QuasiMeasurePreserving f (c • μa) (c • μb) := - ⟨hf.1, by rw [Measure.map_smul]; exact hf.2.smul c⟩ - -theorem ae_map_le (h : QuasiMeasurePreserving f μa μb) : ae (μa.map f) ≤ ae μb := - h.2.ae_le - -theorem tendsto_ae (h : QuasiMeasurePreserving f μa μb) : Tendsto f (ae μa) (ae μb) := - (tendsto_ae_map h.aemeasurable).mono_right h.ae_map_le - -theorem ae (h : QuasiMeasurePreserving f μa μb) {p : β → Prop} (hg : ∀ᵐ x ∂μb, p x) : - ∀ᵐ x ∂μa, p (f x) := - h.tendsto_ae hg - -theorem ae_eq (h : QuasiMeasurePreserving f μa μb) {g₁ g₂ : β → δ} (hg : g₁ =ᵐ[μb] g₂) : - g₁ ∘ f =ᵐ[μa] g₂ ∘ f := - h.ae hg - -theorem preimage_null (h : QuasiMeasurePreserving f μa μb) {s : Set β} (hs : μb s = 0) : - μa (f ⁻¹' s) = 0 := - preimage_null_of_map_null h.aemeasurable (h.2 hs) - -theorem preimage_mono_ae {s t : Set β} (hf : QuasiMeasurePreserving f μa μb) (h : s ≤ᵐ[μb] t) : - f ⁻¹' s ≤ᵐ[μa] f ⁻¹' t := - eventually_map.mp <| - Eventually.filter_mono (tendsto_ae_map hf.aemeasurable) (Eventually.filter_mono hf.ae_map_le h) - -theorem preimage_ae_eq {s t : Set β} (hf : QuasiMeasurePreserving f μa μb) (h : s =ᵐ[μb] t) : - f ⁻¹' s =ᵐ[μa] f ⁻¹' t := - EventuallyLE.antisymm (hf.preimage_mono_ae h.le) (hf.preimage_mono_ae h.symm.le) - -/-- The preimage of a null measurable set under a (quasi) measure preserving map is a null -measurable set. -/ -theorem _root_.MeasureTheory.NullMeasurableSet.preimage {s : Set β} (hs : NullMeasurableSet s μb) - (hf : QuasiMeasurePreserving f μa μb) : NullMeasurableSet (f ⁻¹' s) μa := - let ⟨t, htm, hst⟩ := hs - ⟨f ⁻¹' t, hf.measurable htm, hf.preimage_ae_eq hst⟩ - -theorem preimage_iterate_ae_eq {s : Set α} {f : α → α} (hf : QuasiMeasurePreserving f μ μ) (k : ℕ) - (hs : f ⁻¹' s =ᵐ[μ] s) : f^[k] ⁻¹' s =ᵐ[μ] s := by - induction' k with k ih; · rfl - rw [iterate_succ, preimage_comp] - exact EventuallyEq.trans (hf.preimage_ae_eq ih) hs - -theorem image_zpow_ae_eq {s : Set α} {e : α ≃ α} (he : QuasiMeasurePreserving e μ μ) - (he' : QuasiMeasurePreserving e.symm μ μ) (k : ℤ) (hs : e '' s =ᵐ[μ] s) : - (⇑(e ^ k)) '' s =ᵐ[μ] s := by - rw [Equiv.image_eq_preimage] - obtain ⟨k, rfl | rfl⟩ := k.eq_nat_or_neg - · replace hs : (⇑e⁻¹) ⁻¹' s =ᵐ[μ] s := by rwa [Equiv.image_eq_preimage] at hs - replace he' : (⇑e⁻¹)^[k] ⁻¹' s =ᵐ[μ] s := he'.preimage_iterate_ae_eq k hs - rwa [Equiv.Perm.iterate_eq_pow e⁻¹ k, inv_pow e k] at he' - · rw [zpow_neg, zpow_natCast] - replace hs : e ⁻¹' s =ᵐ[μ] s := by - convert he.preimage_ae_eq hs.symm - rw [Equiv.preimage_image] - replace he : (⇑e)^[k] ⁻¹' s =ᵐ[μ] s := he.preimage_iterate_ae_eq k hs - rwa [Equiv.Perm.iterate_eq_pow e k] at he - --- Need to specify `α := Set α` below because of diamond; see https://github.com/leanprover-community/mathlib4/issues/10941 -theorem limsup_preimage_iterate_ae_eq {f : α → α} (hf : QuasiMeasurePreserving f μ μ) - (hs : f ⁻¹' s =ᵐ[μ] s) : limsup (α := Set α) (fun n => (preimage f)^[n] s) atTop =ᵐ[μ] s := - limsup_ae_eq_of_forall_ae_eq (fun n => (preimage f)^[n] s) fun n ↦ by - simpa only [Set.preimage_iterate_eq] using hf.preimage_iterate_ae_eq n hs - --- Need to specify `α := Set α` below because of diamond; see https://github.com/leanprover-community/mathlib4/issues/10941 -theorem liminf_preimage_iterate_ae_eq {f : α → α} (hf : QuasiMeasurePreserving f μ μ) - (hs : f ⁻¹' s =ᵐ[μ] s) : liminf (α := Set α) (fun n => (preimage f)^[n] s) atTop =ᵐ[μ] s := - liminf_ae_eq_of_forall_ae_eq (fun n => (preimage f)^[n] s) fun n ↦ by - simpa only [Set.preimage_iterate_eq] using hf.preimage_iterate_ae_eq n hs - -/-- For a quasi measure preserving self-map `f`, if a null measurable set `s` is a.e. invariant, -then it is a.e. equal to a measurable invariant set. --/ -theorem exists_preimage_eq_of_preimage_ae {f : α → α} (h : QuasiMeasurePreserving f μ μ) - (hs : NullMeasurableSet s μ) (hs' : f ⁻¹' s =ᵐ[μ] s) : - ∃ t : Set α, MeasurableSet t ∧ t =ᵐ[μ] s ∧ f ⁻¹' t = t := by - obtain ⟨t, htm, ht⟩ := hs - refine ⟨limsup (f^[·] ⁻¹' t) atTop, ?_, ?_, ?_⟩ - · exact .measurableSet_limsup fun n ↦ h.measurable.iterate n htm - · have : f ⁻¹' t =ᵐ[μ] t := (h.preimage_ae_eq ht.symm).trans (hs'.trans ht) - exact limsup_ae_eq_of_forall_ae_eq _ fun n ↦ .trans (h.preimage_iterate_ae_eq _ this) ht.symm - · simp only [Set.preimage_iterate_eq] - exact CompleteLatticeHom.apply_limsup_iterate (CompleteLatticeHom.setPreimage f) t - -open Pointwise - -@[to_additive] -theorem smul_ae_eq_of_ae_eq {G α : Type*} [Group G] [MulAction G α] {_ : MeasurableSpace α} - {s t : Set α} {μ : Measure α} (g : G) - (h_qmp : QuasiMeasurePreserving (g⁻¹ • · : α → α) μ μ) - (h_ae_eq : s =ᵐ[μ] t) : (g • s : Set α) =ᵐ[μ] (g • t : Set α) := by - simpa only [← preimage_smul_inv] using h_qmp.ae_eq h_ae_eq - -end QuasiMeasurePreserving - -section Pointwise - -open Pointwise - -@[to_additive] -theorem pairwise_aedisjoint_of_aedisjoint_forall_ne_one {G α : Type*} [Group G] [MulAction G α] - {_ : MeasurableSpace α} {μ : Measure α} {s : Set α} - (h_ae_disjoint : ∀ g ≠ (1 : G), AEDisjoint μ (g • s) s) - (h_qmp : ∀ g : G, QuasiMeasurePreserving (g • ·) μ μ) : - Pairwise (AEDisjoint μ on fun g : G => g • s) := by - intro g₁ g₂ hg - let g := g₂⁻¹ * g₁ - replace hg : g ≠ 1 := by - rw [Ne, inv_mul_eq_one] - exact hg.symm - have : (g₂⁻¹ • ·) ⁻¹' (g • s ∩ s) = g₁ • s ∩ g₂ • s := by - rw [preimage_eq_iff_eq_image (MulAction.bijective g₂⁻¹), image_smul, smul_set_inter, smul_smul, - smul_smul, inv_mul_cancel, one_smul] - change μ (g₁ • s ∩ g₂ • s) = 0 - exact this ▸ (h_qmp g₂⁻¹).preimage_null (h_ae_disjoint g hg) - -end Pointwise - /-! ### The `cofinite` filter -/ /-- The filter of sets `s` such that `sᶜ` has finite measure. -/ @@ -1785,17 +1305,6 @@ lemma _root_.AEMeasurable.nullMeasurableSet_preimage {f : α → β} {s : Set β (hf : AEMeasurable f μ) (hs : MeasurableSet s) : NullMeasurableSet (f ⁻¹' s) μ := hf.nullMeasurable hs -theorem NullMeasurableSet.mono_ac (h : NullMeasurableSet s μ) (hle : ν ≪ μ) : - NullMeasurableSet s ν := - h.preimage <| (QuasiMeasurePreserving.id μ).mono_left hle - -theorem NullMeasurableSet.mono (h : NullMeasurableSet s μ) (hle : ν ≤ μ) : NullMeasurableSet s ν := - h.mono_ac hle.absolutelyContinuous - -theorem AEDisjoint.preimage {ν : Measure β} {f : α → β} {s t : Set β} (ht : AEDisjoint ν s t) - (hf : QuasiMeasurePreserving f μ ν) : AEDisjoint μ (f ⁻¹' s) (f ⁻¹' t) := - hf.preimage_null ht - @[simp] theorem ae_eq_bot : ae μ = ⊥ ↔ μ = 0 := by rw [← empty_mem_iff_bot, mem_ae_iff, compl_empty, measure_univ_eq_zero] @@ -1810,35 +1319,6 @@ instance Measure.ae.neBot [NeZero μ] : (ae μ).NeBot := ae_neBot.2 <| NeZero.ne theorem ae_zero {_m0 : MeasurableSpace α} : ae (0 : Measure α) = ⊥ := ae_eq_bot.2 rfl -@[mono] -theorem ae_mono (h : μ ≤ ν) : ae μ ≤ ae ν := - h.absolutelyContinuous.ae_le - -theorem mem_ae_map_iff {f : α → β} (hf : AEMeasurable f μ) {s : Set β} (hs : MeasurableSet s) : - s ∈ ae (μ.map f) ↔ f ⁻¹' s ∈ ae μ := by - simp only [mem_ae_iff, map_apply_of_aemeasurable hf hs.compl, preimage_compl] - -theorem mem_ae_of_mem_ae_map {f : α → β} (hf : AEMeasurable f μ) {s : Set β} - (hs : s ∈ ae (μ.map f)) : f ⁻¹' s ∈ ae μ := - (tendsto_ae_map hf).eventually hs - -theorem ae_map_iff {f : α → β} (hf : AEMeasurable f μ) {p : β → Prop} - (hp : MeasurableSet { x | p x }) : (∀ᵐ y ∂μ.map f, p y) ↔ ∀ᵐ x ∂μ, p (f x) := - mem_ae_map_iff hf hp - -theorem ae_of_ae_map {f : α → β} (hf : AEMeasurable f μ) {p : β → Prop} (h : ∀ᵐ y ∂μ.map f, p y) : - ∀ᵐ x ∂μ, p (f x) := - mem_ae_of_mem_ae_map hf h - -theorem ae_map_mem_range {m0 : MeasurableSpace α} (f : α → β) (hf : MeasurableSet (range f)) - (μ : Measure α) : ∀ᵐ x ∂μ.map f, x ∈ range f := by - by_cases h : AEMeasurable f μ - · change range f ∈ ae (μ.map f) - rw [mem_ae_map_iff h hf] - filter_upwards using mem_range_self - · simp [map_of_not_aemeasurable h] - - section Intervals theorem biSup_measure_Iic [Preorder α] {s : Set α} (hsc : s.Countable) @@ -1903,76 +1383,4 @@ end end MeasureTheory -namespace MeasurableEmbedding - -open MeasureTheory Measure - -variable {m0 : MeasurableSpace α} {m1 : MeasurableSpace β} {f : α → β} {μ ν : Measure α} - -nonrec theorem map_apply (hf : MeasurableEmbedding f) (μ : Measure α) (s : Set β) : - μ.map f s = μ (f ⁻¹' s) := by - refine le_antisymm ?_ (le_map_apply hf.measurable.aemeasurable s) - set t := f '' toMeasurable μ (f ⁻¹' s) ∪ (range f)ᶜ - have htm : MeasurableSet t := - (hf.measurableSet_image.2 <| measurableSet_toMeasurable _ _).union - hf.measurableSet_range.compl - have hst : s ⊆ t := by - rw [subset_union_compl_iff_inter_subset, ← image_preimage_eq_inter_range] - exact image_subset _ (subset_toMeasurable _ _) - have hft : f ⁻¹' t = toMeasurable μ (f ⁻¹' s) := by - rw [preimage_union, preimage_compl, preimage_range, compl_univ, union_empty, - hf.injective.preimage_image] - calc - μ.map f s ≤ μ.map f t := by gcongr - _ = μ (f ⁻¹' s) := by rw [map_apply hf.measurable htm, hft, measure_toMeasurable] - -lemma absolutelyContinuous_map (hf : MeasurableEmbedding f) (hμν : μ ≪ ν) : - μ.map f ≪ ν.map f := by - intro t ht - rw [hf.map_apply] at ht ⊢ - exact hμν ht - -end MeasurableEmbedding - -namespace MeasurableEquiv - -/-! Interactions of measurable equivalences and measures -/ - -open Equiv MeasureTheory.Measure - -variable {_ : MeasurableSpace α} [MeasurableSpace β] {μ : Measure α} {ν : Measure β} - -/-- If we map a measure along a measurable equivalence, we can compute the measure on all sets - (not just the measurable ones). -/ -protected theorem map_apply (f : α ≃ᵐ β) (s : Set β) : μ.map f s = μ (f ⁻¹' s) := - f.measurableEmbedding.map_apply _ _ - -@[simp] -theorem map_symm_map (e : α ≃ᵐ β) : (μ.map e).map e.symm = μ := by - simp [map_map e.symm.measurable e.measurable] - -@[simp] -theorem map_map_symm (e : α ≃ᵐ β) : (ν.map e.symm).map e = ν := by - simp [map_map e.measurable e.symm.measurable] - -theorem map_measurableEquiv_injective (e : α ≃ᵐ β) : Injective (Measure.map e) := by - intro μ₁ μ₂ hμ - apply_fun Measure.map e.symm at hμ - simpa [map_symm_map e] using hμ - -theorem map_apply_eq_iff_map_symm_apply_eq (e : α ≃ᵐ β) : μ.map e = ν ↔ ν.map e.symm = μ := by - rw [← (map_measurableEquiv_injective e).eq_iff, map_map_symm, eq_comm] - -theorem map_ae (f : α ≃ᵐ β) (μ : Measure α) : Filter.map f (ae μ) = ae (map f μ) := by - ext s - simp_rw [mem_map, mem_ae_iff, ← preimage_compl, f.map_apply] - -theorem quasiMeasurePreserving_symm (μ : Measure α) (e : α ≃ᵐ β) : - QuasiMeasurePreserving e.symm (map e μ) μ := - ⟨e.symm.measurable, by rw [Measure.map_map, e.symm_comp_self, Measure.map_id] <;> measurability⟩ - -end MeasurableEquiv - end - -set_option linter.style.longFile 2100 diff --git a/Mathlib/MeasureTheory/Measure/NullMeasurable.lean b/Mathlib/MeasureTheory/Measure/NullMeasurable.lean index b2e786979552e..0e4f0916c5455 100644 --- a/Mathlib/MeasureTheory/Measure/NullMeasurable.lean +++ b/Mathlib/MeasureTheory/Measure/NullMeasurable.lean @@ -207,6 +207,8 @@ end NullMeasurableSet open NullMeasurableSet +open scoped Function -- required for scoped `on` notation + /-- If `sᵢ` is a countable family of (null) measurable pairwise `μ`-a.e. disjoint sets, then there exists a subordinate family `tᵢ ⊆ sᵢ` of measurable pairwise disjoint sets such that `tᵢ =ᵐ[μ] sᵢ`. -/ diff --git a/Mathlib/MeasureTheory/Measure/Prod.lean b/Mathlib/MeasureTheory/Measure/Prod.lean index ce852342469e9..951fd5c3bf2c4 100644 --- a/Mathlib/MeasureTheory/Measure/Prod.lean +++ b/Mathlib/MeasureTheory/Measure/Prod.lean @@ -308,7 +308,7 @@ theorem measure_ae_null_of_prod_null {s : Set (α × β)} (h : μ.prod ν s = 0) rw [measure_prod_null mt] at ht rw [eventuallyLE_antisymm_iff] exact - ⟨EventuallyLE.trans_eq (Eventually.of_forall fun x => (measure_mono (preimage_mono hst) : _)) + ⟨EventuallyLE.trans_eq (Eventually.of_forall fun x => measure_mono (preimage_mono hst)) ht, Eventually.of_forall fun x => zero_le _⟩ diff --git a/Mathlib/MeasureTheory/Measure/QuasiMeasurePreserving.lean b/Mathlib/MeasureTheory/Measure/QuasiMeasurePreserving.lean new file mode 100644 index 0000000000000..0630e9e86c598 --- /dev/null +++ b/Mathlib/MeasureTheory/Measure/QuasiMeasurePreserving.lean @@ -0,0 +1,229 @@ +/- +Copyright (c) 2017 Johannes Hölzl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johannes Hölzl, Mario Carneiro +-/ +import Mathlib.MeasureTheory.Measure.AbsolutelyContinuous +import Mathlib.MeasureTheory.OuterMeasure.BorelCantelli + +/-! +# Quasi Measure Preserving Functions + +A map `f : α → β` is said to be *quasi measure preserving* (a.k.a. non-singular) w.r.t. measures +`μa` and `μb` if it is measurable and `μb s = 0` implies `μa (f ⁻¹' s) = 0`. +That last condition can also be written `μa.map f ≪ μb` (the map of `μa` by `f` is +absolutely continuous with respect to `μb`). + +## Main definitions + +* `MeasureTheory.Measure.QuasiMeasurePreserving f μa μb`: `f` is quasi measure preserving with + respect to `μa` and `μb`. + +-/ + +variable {α β γ δ : Type*} + +namespace MeasureTheory + +open Set Function ENNReal +open Filter hiding map + +variable {mα : MeasurableSpace α} {mβ : MeasurableSpace β} {mγ : MeasurableSpace γ} + {μ ν : Measure α} {s : Set α} + +namespace Measure + +/-- A map `f : α → β` is said to be *quasi measure preserving* (a.k.a. non-singular) w.r.t. measures +`μa` and `μb` if it is measurable and `μb s = 0` implies `μa (f ⁻¹' s) = 0`. -/ +structure QuasiMeasurePreserving {m0 : MeasurableSpace α} (f : α → β) + (μa : Measure α := by volume_tac) + (μb : Measure β := by volume_tac) : Prop where + protected measurable : Measurable f + protected absolutelyContinuous : μa.map f ≪ μb + +namespace QuasiMeasurePreserving + +protected theorem id {_m0 : MeasurableSpace α} (μ : Measure α) : QuasiMeasurePreserving id μ μ := + ⟨measurable_id, map_id.absolutelyContinuous⟩ + +variable {μa μa' : Measure α} {μb μb' : Measure β} {μc : Measure γ} {f : α → β} + +protected theorem _root_.Measurable.quasiMeasurePreserving + {_m0 : MeasurableSpace α} (hf : Measurable f) (μ : Measure α) : + QuasiMeasurePreserving f μ (μ.map f) := + ⟨hf, AbsolutelyContinuous.rfl⟩ + +theorem mono_left (h : QuasiMeasurePreserving f μa μb) (ha : μa' ≪ μa) : + QuasiMeasurePreserving f μa' μb := + ⟨h.1, (ha.map h.1).trans h.2⟩ + +theorem mono_right (h : QuasiMeasurePreserving f μa μb) (ha : μb ≪ μb') : + QuasiMeasurePreserving f μa μb' := + ⟨h.1, h.2.trans ha⟩ + +@[mono] +theorem mono (ha : μa' ≪ μa) (hb : μb ≪ μb') (h : QuasiMeasurePreserving f μa μb) : + QuasiMeasurePreserving f μa' μb' := + (h.mono_left ha).mono_right hb + +protected theorem comp {g : β → γ} {f : α → β} (hg : QuasiMeasurePreserving g μb μc) + (hf : QuasiMeasurePreserving f μa μb) : QuasiMeasurePreserving (g ∘ f) μa μc := + ⟨hg.measurable.comp hf.measurable, by + rw [← map_map hg.1 hf.1] + exact (hf.2.map hg.1).trans hg.2⟩ + +protected theorem iterate {f : α → α} (hf : QuasiMeasurePreserving f μa μa) : + ∀ n, QuasiMeasurePreserving f^[n] μa μa + | 0 => QuasiMeasurePreserving.id μa + | n + 1 => (hf.iterate n).comp hf + +protected theorem aemeasurable (hf : QuasiMeasurePreserving f μa μb) : AEMeasurable f μa := + hf.1.aemeasurable + +theorem smul_measure {R : Type*} [SMul R ℝ≥0∞] [IsScalarTower R ℝ≥0∞ ℝ≥0∞] + (hf : QuasiMeasurePreserving f μa μb) (c : R) : QuasiMeasurePreserving f (c • μa) (c • μb) := + ⟨hf.1, by rw [Measure.map_smul]; exact hf.2.smul c⟩ + +theorem ae_map_le (h : QuasiMeasurePreserving f μa μb) : ae (μa.map f) ≤ ae μb := + h.2.ae_le + +theorem tendsto_ae (h : QuasiMeasurePreserving f μa μb) : Tendsto f (ae μa) (ae μb) := + (tendsto_ae_map h.aemeasurable).mono_right h.ae_map_le + +theorem ae (h : QuasiMeasurePreserving f μa μb) {p : β → Prop} (hg : ∀ᵐ x ∂μb, p x) : + ∀ᵐ x ∂μa, p (f x) := + h.tendsto_ae hg + +theorem ae_eq (h : QuasiMeasurePreserving f μa μb) {g₁ g₂ : β → δ} (hg : g₁ =ᵐ[μb] g₂) : + g₁ ∘ f =ᵐ[μa] g₂ ∘ f := + h.ae hg + +theorem preimage_null (h : QuasiMeasurePreserving f μa μb) {s : Set β} (hs : μb s = 0) : + μa (f ⁻¹' s) = 0 := + preimage_null_of_map_null h.aemeasurable (h.2 hs) + +theorem preimage_mono_ae {s t : Set β} (hf : QuasiMeasurePreserving f μa μb) (h : s ≤ᵐ[μb] t) : + f ⁻¹' s ≤ᵐ[μa] f ⁻¹' t := + eventually_map.mp <| + Eventually.filter_mono (tendsto_ae_map hf.aemeasurable) (Eventually.filter_mono hf.ae_map_le h) + +theorem preimage_ae_eq {s t : Set β} (hf : QuasiMeasurePreserving f μa μb) (h : s =ᵐ[μb] t) : + f ⁻¹' s =ᵐ[μa] f ⁻¹' t := + EventuallyLE.antisymm (hf.preimage_mono_ae h.le) (hf.preimage_mono_ae h.symm.le) + +/-- The preimage of a null measurable set under a (quasi) measure preserving map is a null +measurable set. -/ +theorem _root_.MeasureTheory.NullMeasurableSet.preimage {s : Set β} (hs : NullMeasurableSet s μb) + (hf : QuasiMeasurePreserving f μa μb) : NullMeasurableSet (f ⁻¹' s) μa := + let ⟨t, htm, hst⟩ := hs + ⟨f ⁻¹' t, hf.measurable htm, hf.preimage_ae_eq hst⟩ + +theorem preimage_iterate_ae_eq {s : Set α} {f : α → α} (hf : QuasiMeasurePreserving f μ μ) (k : ℕ) + (hs : f ⁻¹' s =ᵐ[μ] s) : f^[k] ⁻¹' s =ᵐ[μ] s := by + induction' k with k ih; · rfl + rw [iterate_succ, preimage_comp] + exact EventuallyEq.trans (hf.preimage_ae_eq ih) hs + +theorem image_zpow_ae_eq {s : Set α} {e : α ≃ α} (he : QuasiMeasurePreserving e μ μ) + (he' : QuasiMeasurePreserving e.symm μ μ) (k : ℤ) (hs : e '' s =ᵐ[μ] s) : + (⇑(e ^ k)) '' s =ᵐ[μ] s := by + rw [Equiv.image_eq_preimage] + obtain ⟨k, rfl | rfl⟩ := k.eq_nat_or_neg + · replace hs : (⇑e⁻¹) ⁻¹' s =ᵐ[μ] s := by rwa [Equiv.image_eq_preimage] at hs + replace he' : (⇑e⁻¹)^[k] ⁻¹' s =ᵐ[μ] s := he'.preimage_iterate_ae_eq k hs + rwa [Equiv.Perm.iterate_eq_pow e⁻¹ k, inv_pow e k] at he' + · rw [zpow_neg, zpow_natCast] + replace hs : e ⁻¹' s =ᵐ[μ] s := by + convert he.preimage_ae_eq hs.symm + rw [Equiv.preimage_image] + replace he : (⇑e)^[k] ⁻¹' s =ᵐ[μ] s := he.preimage_iterate_ae_eq k hs + rwa [Equiv.Perm.iterate_eq_pow e k] at he + +-- Need to specify `α := Set α` below because of diamond; see https://github.com/leanprover-community/mathlib4/issues/10941 +theorem limsup_preimage_iterate_ae_eq {f : α → α} (hf : QuasiMeasurePreserving f μ μ) + (hs : f ⁻¹' s =ᵐ[μ] s) : limsup (α := Set α) (fun n => (preimage f)^[n] s) atTop =ᵐ[μ] s := + limsup_ae_eq_of_forall_ae_eq (fun n => (preimage f)^[n] s) fun n ↦ by + simpa only [Set.preimage_iterate_eq] using hf.preimage_iterate_ae_eq n hs + +-- Need to specify `α := Set α` below because of diamond; see https://github.com/leanprover-community/mathlib4/issues/10941 +theorem liminf_preimage_iterate_ae_eq {f : α → α} (hf : QuasiMeasurePreserving f μ μ) + (hs : f ⁻¹' s =ᵐ[μ] s) : liminf (α := Set α) (fun n => (preimage f)^[n] s) atTop =ᵐ[μ] s := + liminf_ae_eq_of_forall_ae_eq (fun n => (preimage f)^[n] s) fun n ↦ by + simpa only [Set.preimage_iterate_eq] using hf.preimage_iterate_ae_eq n hs + +/-- For a quasi measure preserving self-map `f`, if a null measurable set `s` is a.e. invariant, +then it is a.e. equal to a measurable invariant set. +-/ +theorem exists_preimage_eq_of_preimage_ae {f : α → α} (h : QuasiMeasurePreserving f μ μ) + (hs : NullMeasurableSet s μ) (hs' : f ⁻¹' s =ᵐ[μ] s) : + ∃ t : Set α, MeasurableSet t ∧ t =ᵐ[μ] s ∧ f ⁻¹' t = t := by + obtain ⟨t, htm, ht⟩ := hs + refine ⟨limsup (f^[·] ⁻¹' t) atTop, ?_, ?_, ?_⟩ + · exact .measurableSet_limsup fun n ↦ h.measurable.iterate n htm + · have : f ⁻¹' t =ᵐ[μ] t := (h.preimage_ae_eq ht.symm).trans (hs'.trans ht) + exact limsup_ae_eq_of_forall_ae_eq _ fun n ↦ .trans (h.preimage_iterate_ae_eq _ this) ht.symm + · simp only [Set.preimage_iterate_eq] + exact CompleteLatticeHom.apply_limsup_iterate (CompleteLatticeHom.setPreimage f) t + +open Pointwise + +@[to_additive] +theorem smul_ae_eq_of_ae_eq {G α : Type*} [Group G] [MulAction G α] {_ : MeasurableSpace α} + {s t : Set α} {μ : Measure α} (g : G) + (h_qmp : QuasiMeasurePreserving (g⁻¹ • · : α → α) μ μ) + (h_ae_eq : s =ᵐ[μ] t) : (g • s : Set α) =ᵐ[μ] (g • t : Set α) := by + simpa only [← preimage_smul_inv] using h_qmp.ae_eq h_ae_eq + +end QuasiMeasurePreserving + +section Pointwise + +open Pointwise + +@[to_additive] +theorem pairwise_aedisjoint_of_aedisjoint_forall_ne_one {G α : Type*} [Group G] [MulAction G α] + {_ : MeasurableSpace α} {μ : Measure α} {s : Set α} + (h_ae_disjoint : ∀ g ≠ (1 : G), AEDisjoint μ (g • s) s) + (h_qmp : ∀ g : G, QuasiMeasurePreserving (g • ·) μ μ) : + Pairwise (AEDisjoint μ on fun g : G => g • s) := by + intro g₁ g₂ hg + let g := g₂⁻¹ * g₁ + replace hg : g ≠ 1 := by + rw [Ne, inv_mul_eq_one] + exact hg.symm + have : (g₂⁻¹ • ·) ⁻¹' (g • s ∩ s) = g₁ • s ∩ g₂ • s := by + rw [preimage_eq_iff_eq_image (MulAction.bijective g₂⁻¹), image_smul, smul_set_inter, smul_smul, + smul_smul, inv_mul_cancel, one_smul] + change μ (g₁ • s ∩ g₂ • s) = 0 + exact this ▸ (h_qmp g₂⁻¹).preimage_null (h_ae_disjoint g hg) + +end Pointwise + +end Measure + +open Measure + +theorem NullMeasurableSet.mono_ac (h : NullMeasurableSet s μ) (hle : ν ≪ μ) : + NullMeasurableSet s ν := + h.preimage <| (QuasiMeasurePreserving.id μ).mono_left hle + +theorem NullMeasurableSet.mono (h : NullMeasurableSet s μ) (hle : ν ≤ μ) : NullMeasurableSet s ν := + h.mono_ac hle.absolutelyContinuous + +theorem AEDisjoint.preimage {ν : Measure β} {f : α → β} {s t : Set β} (ht : AEDisjoint ν s t) + (hf : QuasiMeasurePreserving f μ ν) : AEDisjoint μ (f ⁻¹' s) (f ⁻¹' t) := + hf.preimage_null ht + +end MeasureTheory + +open MeasureTheory + +namespace MeasurableEquiv + +variable {_ : MeasurableSpace α} [MeasurableSpace β] {μ : Measure α} {ν : Measure β} + +theorem quasiMeasurePreserving_symm (μ : Measure α) (e : α ≃ᵐ β) : + Measure.QuasiMeasurePreserving e.symm (μ.map e) μ := + ⟨e.symm.measurable, by rw [Measure.map_map, e.symm_comp_self, Measure.map_id] <;> measurability⟩ + +end MeasurableEquiv diff --git a/Mathlib/MeasureTheory/Measure/Regular.lean b/Mathlib/MeasureTheory/Measure/Regular.lean index 91b924e0a4f30..815cd81e7c71e 100644 --- a/Mathlib/MeasureTheory/Measure/Regular.lean +++ b/Mathlib/MeasureTheory/Measure/Regular.lean @@ -386,6 +386,7 @@ instance smul_nnreal (μ : Measure α) [OuterRegular μ] (c : ℝ≥0) : OuterRegular (c • μ) := OuterRegular.smul μ coe_ne_top +open scoped Function in -- required for scoped `on` notation /-- If the restrictions of a measure to countably many open sets covering the space are outer regular, then the measure itself is outer regular. -/ lemma of_restrict [OpensMeasurableSpace α] {μ : Measure α} {s : ℕ → Set α} diff --git a/Mathlib/MeasureTheory/Measure/RegularityCompacts.lean b/Mathlib/MeasureTheory/Measure/RegularityCompacts.lean new file mode 100644 index 0000000000000..1d8e91a4aa967 --- /dev/null +++ b/Mathlib/MeasureTheory/Measure/RegularityCompacts.lean @@ -0,0 +1,248 @@ +/- +Copyright (c) 2023 Rémy Degenne. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rémy Degenne, Peter Pfaffelhuber +-/ +import Mathlib.Analysis.SpecificLimits.Basic +import Mathlib.MeasureTheory.Measure.Regular +import Mathlib.Topology.MetricSpace.Polish +import Mathlib.Topology.UniformSpace.Cauchy + +/-! +# Inner regularity of finite measures + +The main result of this file is `theorem InnerRegularCompactLTTop`: +A finite measure `μ` on a `PseudoEMetricSpace E` and `CompleteSpace E` with +`SecondCountableTopology E` is inner regular with respect to compact sets. In other +words, a finite measure on such a space is a tight measure. + +Finite measures on Polish spaces are an important special case, which makes the result +`theorem PolishSpace.innerRegular_isCompact_measurableSet` an important result in probability. +-/ + +open Set MeasureTheory + +open scoped ENNReal + +namespace MeasureTheory + +variable {α : Type*} [MeasurableSpace α] {μ : Measure α} + +theorem innerRegularWRT_isCompact_closure_iff [TopologicalSpace α] [R1Space α] : + μ.InnerRegularWRT (IsCompact ∘ closure) IsClosed ↔ μ.InnerRegularWRT IsCompact IsClosed := by + constructor <;> intro h A hA r hr + · rcases h hA r hr with ⟨K, ⟨hK1, hK2, hK3⟩⟩ + exact ⟨closure K, closure_minimal hK1 hA, hK2, hK3.trans_le (measure_mono subset_closure)⟩ + · rcases h hA r hr with ⟨K, ⟨hK1, hK2, hK3⟩⟩ + refine ⟨closure K, closure_minimal hK1 hA, ?_, ?_⟩ + · simpa only [closure_closure, Function.comp_apply] using hK2.closure + · exact hK3.trans_le (measure_mono subset_closure) + +lemma innerRegularWRT_isCompact_isClosed_iff_innerRegularWRT_isCompact_closure + [TopologicalSpace α] [R1Space α] : + μ.InnerRegularWRT (fun s ↦ IsCompact s ∧ IsClosed s) IsClosed + ↔ μ.InnerRegularWRT (IsCompact ∘ closure) IsClosed := by + constructor <;> intro h A hA r hr + · obtain ⟨K, hK1, ⟨hK2, _⟩, hK4⟩ := h hA r hr + refine ⟨K, hK1, ?_, hK4⟩ + simp only [closure_closure, Function.comp_apply] + exact hK2.closure + · obtain ⟨K, hK1, hK2, hK3⟩ := h hA r hr + refine ⟨closure K, closure_minimal hK1 hA, ?_, ?_⟩ + · simpa only [isClosed_closure, and_true] + · exact hK3.trans_le (measure_mono subset_closure) + +lemma innerRegularWRT_isCompact_isClosed_iff [TopologicalSpace α] [R1Space α] : + μ.InnerRegularWRT (fun s ↦ IsCompact s ∧ IsClosed s) IsClosed + ↔ μ.InnerRegularWRT IsCompact IsClosed := + innerRegularWRT_isCompact_isClosed_iff_innerRegularWRT_isCompact_closure.trans + innerRegularWRT_isCompact_closure_iff + +/-- +If predicate `p` is preserved under intersections with sets satisfying predicate `q`, and sets +satisfying `p` cover the space arbitrarily well, then `μ` is inner regular with respect to +predicates `p` and `q`. +-/ +theorem innerRegularWRT_of_exists_compl_lt {p q : Set α → Prop} (hpq : ∀ A B, p A → q B → p (A ∩ B)) + (hμ : ∀ ε, 0 < ε → ∃ K, p K ∧ μ Kᶜ < ε) : + μ.InnerRegularWRT p q := by + intro A hA r hr + obtain ⟨K, hK, hK_subset, h_lt⟩ : ∃ K, p K ∧ K ⊆ A ∧ μ (A \ K) < μ A - r := by + obtain ⟨K', hpK', hK'_lt⟩ := hμ (μ A - r) (tsub_pos_of_lt hr) + refine ⟨K' ∩ A, hpq K' A hpK' hA, inter_subset_right, ?_⟩ + · refine (measure_mono fun x ↦ ?_).trans_lt hK'_lt + simp only [diff_inter_self_eq_diff, mem_diff, mem_compl_iff, and_imp, imp_self, imp_true_iff] + refine ⟨K, hK_subset, hK, ?_⟩ + have h_lt' : μ A - μ K < μ A - r := le_measure_diff.trans_lt h_lt + exact lt_of_tsub_lt_tsub_left h_lt' + +theorem innerRegularWRT_isCompact_closure_of_univ [TopologicalSpace α] + (hμ : ∀ ε, 0 < ε → ∃ K, IsCompact (closure K) ∧ μ (Kᶜ) < ε) : + μ.InnerRegularWRT (IsCompact ∘ closure) IsClosed := by + refine innerRegularWRT_of_exists_compl_lt (fun s t hs ht ↦ ?_) hμ + have : IsCompact (closure s ∩ t) := hs.inter_right ht + refine this.of_isClosed_subset isClosed_closure ?_ + refine (closure_inter_subset_inter_closure _ _).trans_eq ?_ + rw [IsClosed.closure_eq ht] + +theorem exists_isCompact_closure_measure_compl_lt [UniformSpace α] [CompleteSpace α] + [SecondCountableTopology α] [(uniformity α).IsCountablyGenerated] + [OpensMeasurableSpace α] (P : Measure α) [IsFiniteMeasure P] (ε : ℝ≥0∞) (hε : 0 < ε) : + ∃ K, IsCompact (closure K) ∧ P Kᶜ < ε := by + /- + If α is empty, the result is trivial. + + Otherwise, fix a dense sequence `seq` and an antitone basis `t` of entourages. We find a sequence + of natural numbers `u n`, such that `interUnionBalls seq u t`, which is the intersection over + `n` of the `t n`-neighborhood of `seq 1, ..., seq (u n)`, covers the space arbitrarily well. + -/ + cases isEmpty_or_nonempty α + case inl => + refine ⟨∅, by simp, ?_⟩ + rw [← Set.univ_eq_empty_iff.mpr] + · simpa only [compl_univ, measure_empty, ENNReal.coe_pos] using hε + · assumption + case inr => + let seq := TopologicalSpace.denseSeq α + have hseq_dense : DenseRange seq := TopologicalSpace.denseRange_denseSeq α + obtain ⟨t : ℕ → Set (α × α), + hto : ∀ i, t i ∈ (uniformity α).sets ∧ IsOpen (t i) ∧ SymmetricRel (t i), + h_basis : (uniformity α).HasAntitoneBasis t⟩ := + (@uniformity_hasBasis_open_symmetric α _).exists_antitone_subbasis + let f : ℕ → ℕ → Set α := fun n m ↦ UniformSpace.ball (seq m) (t n) + have h_univ n : (⋃ m, f n m) = univ := hseq_dense.iUnion_uniformity_ball (hto n).1 + have h3 n (ε : ℝ≥0∞) (hε : 0 < ε) : ∃ m, P (⋂ m' ≤ m, (f n m')ᶜ) < ε := by + refine exists_measure_iInter_lt (fun m ↦ ?_) hε ⟨0, measure_ne_top P _⟩ ?_ + · exact (measurable_prod_mk_left (IsOpen.measurableSet (hto n).2.1)).compl.nullMeasurableSet + · rw [← compl_iUnion, h_univ, compl_univ] + choose! s' s'bound using h3 + rcases ENNReal.exists_pos_sum_of_countable' (ne_of_gt hε) ℕ with ⟨δ, hδ1, hδ2⟩ + classical + let u : ℕ → ℕ := fun n ↦ s' n (δ n) + refine ⟨interUnionBalls seq u t, isCompact_closure_interUnionBalls h_basis.toHasBasis seq u, ?_⟩ + rw [interUnionBalls, Set.compl_iInter] + refine ((measure_iUnion_le _).trans ?_).trans_lt hδ2 + refine ENNReal.tsum_le_tsum (fun n ↦ ?_) + have h'' n : Prod.swap ⁻¹' t n = t n := SymmetricRel.eq (hto n).2.2 + simp only [h'', compl_iUnion, ge_iff_le] + exact (s'bound n (δ n) (hδ1 n)).le + +theorem innerRegularWRT_isCompact_closure [UniformSpace α] [CompleteSpace α] + [SecondCountableTopology α] [(uniformity α).IsCountablyGenerated] + [OpensMeasurableSpace α] (P : Measure α) [IsFiniteMeasure P] : + P.InnerRegularWRT (IsCompact ∘ closure) IsClosed := + innerRegularWRT_isCompact_closure_of_univ + (exists_isCompact_closure_measure_compl_lt P) + +theorem innerRegularWRT_isCompact_isClosed [UniformSpace α] [CompleteSpace α] + [SecondCountableTopology α] [(uniformity α).IsCountablyGenerated] + [OpensMeasurableSpace α] (P : Measure α) [IsFiniteMeasure P] : + P.InnerRegularWRT (fun s ↦ IsCompact s ∧ IsClosed s) IsClosed := by + rw [innerRegularWRT_isCompact_isClosed_iff_innerRegularWRT_isCompact_closure] + exact innerRegularWRT_isCompact_closure P + +theorem innerRegularWRT_isCompact [UniformSpace α] [CompleteSpace α] + [SecondCountableTopology α] [(uniformity α).IsCountablyGenerated] + [OpensMeasurableSpace α] (P : Measure α) [IsFiniteMeasure P] : + P.InnerRegularWRT IsCompact IsClosed := by + rw [← innerRegularWRT_isCompact_closure_iff] + exact innerRegularWRT_isCompact_closure P + +theorem innerRegularWRT_isCompact_isClosed_isOpen [PseudoEMetricSpace α] + [CompleteSpace α] [SecondCountableTopology α] [OpensMeasurableSpace α] + (P : Measure α) [IsFiniteMeasure P] : + P.InnerRegularWRT (fun s ↦ IsCompact s ∧ IsClosed s) IsOpen := + (innerRegularWRT_isCompact_isClosed P).trans + (Measure.InnerRegularWRT.of_pseudoMetrizableSpace P) + +theorem innerRegularWRT_isCompact_isOpen [PseudoEMetricSpace α] + [CompleteSpace α] [SecondCountableTopology α] [OpensMeasurableSpace α] + (P : Measure α) [IsFiniteMeasure P] : + P.InnerRegularWRT IsCompact IsOpen := + (innerRegularWRT_isCompact P).trans + (Measure.InnerRegularWRT.of_pseudoMetrizableSpace P) + +/-- +A finite measure `μ` on a `PseudoEMetricSpace E` and `CompleteSpace E` with +`SecondCountableTopology E` is inner regular. In other words, a finite measure +on such a space is a tight measure. +-/ +instance InnerRegular_of_pseudoEMetricSpace_completeSpace_secondCountable [PseudoEMetricSpace α] + [CompleteSpace α] [SecondCountableTopology α] [BorelSpace α] + (P : Measure α) [IsFiniteMeasure P] : + P.InnerRegular := by + refine @Measure.InnerRegularCompactLTTop.instInnerRegularOfSigmaFinite _ _ _ _ + ⟨Measure.InnerRegularWRT.measurableSet_of_isOpen ?_ ?_⟩ _ + · exact innerRegularWRT_isCompact_isOpen P + · exact fun s t hs_compact ht_open ↦ hs_compact.inter_right ht_open.isClosed_compl + +/-- +A special case of `innerRegular_of_pseudoEMetricSpace_completeSpace_secondCountable` for Polish +spaces: A finite measure on a Polish space is a tight measure. +-/ +instance InnerRegular_of_polishSpace [TopologicalSpace α] + [PolishSpace α] [BorelSpace α] (P : Measure α) [IsFiniteMeasure P] : + P.InnerRegular := by + letI := upgradePolishSpace α + exact InnerRegular_of_pseudoEMetricSpace_completeSpace_secondCountable P + +/-- +A measure `μ` on a `PseudoEMetricSpace E` and `CompleteSpace E` with `SecondCountableTopology E` +is inner regular for finite measure sets with respect to compact sets. +-/ +instance InnerRegularCompactLTTop_of_pseudoEMetricSpace_completeSpace_secondCountable + [PseudoEMetricSpace α] [CompleteSpace α] [SecondCountableTopology α] [BorelSpace α] + (μ : Measure α) : + μ.InnerRegularCompactLTTop := by + constructor; intro A ⟨hA1, hA2⟩ r hr + have IRC : Measure.InnerRegularCompactLTTop (μ.restrict A) := by + exact @Measure.InnerRegular.instInnerRegularCompactLTTop _ _ _ _ + (@InnerRegular_of_pseudoEMetricSpace_completeSpace_secondCountable _ _ _ _ _ _ + (μ.restrict A) (@Restrict.isFiniteMeasure _ _ _ μ (fact_iff.mpr hA2.lt_top))) + have hA2' : (μ.restrict A) A ≠ ⊤ := by + rwa [Measure.restrict_apply_self] + have hr' : r < μ.restrict A A := by + rwa [Measure.restrict_apply_self] + obtain ⟨K, ⟨hK1, hK2, hK3⟩⟩ := @MeasurableSet.exists_lt_isCompact_of_ne_top + _ _ (μ.restrict A) _ IRC _ hA1 hA2' r hr' + use K, hK1, hK2 + rwa [Measure.restrict_eq_self μ hK1] at hK3 + +/-- +A special case of `innerRegularCompactLTTop_of_pseudoEMetricSpace_completeSpace_secondCountable` +for Polish spaces: A measure `μ` on a Polish space inner regular for finite measure sets with +respect to compact sets. +-/ +instance InnerRegularCompactLTTop_of_polishSpace + [TopologicalSpace α] [PolishSpace α] [BorelSpace α] (μ : Measure α) : + μ.InnerRegularCompactLTTop := by + letI := upgradePolishSpace α + exact InnerRegularCompactLTTop_of_pseudoEMetricSpace_completeSpace_secondCountable μ + +theorem innerRegular_isCompact_isClosed_measurableSet_of_finite [PseudoEMetricSpace α] + [CompleteSpace α] [SecondCountableTopology α] [BorelSpace α] + (P : Measure α) [IsFiniteMeasure P] : + P.InnerRegularWRT (fun s ↦ IsCompact s ∧ IsClosed s) MeasurableSet := by + suffices P.InnerRegularWRT (fun s ↦ IsCompact s ∧ IsClosed s) + fun s ↦ MeasurableSet s ∧ P s ≠ ∞ by + convert this + simp only [eq_iff_iff, iff_self_and] + exact fun _ ↦ measure_ne_top P _ + refine Measure.InnerRegularWRT.measurableSet_of_isOpen ?_ ?_ + · exact innerRegularWRT_isCompact_isClosed_isOpen P + · rintro s t ⟨hs_compact, hs_closed⟩ ht_open + rw [diff_eq] + exact ⟨hs_compact.inter_right ht_open.isClosed_compl, + hs_closed.inter (isClosed_compl_iff.mpr ht_open)⟩ + +/-- +On a Polish space, any finite measure is regular with respect to compact and closed sets. In +particular, a finite measure on a Polish space is a tight measure. +-/ +theorem PolishSpace.innerRegular_isCompact_isClosed_measurableSet [TopologicalSpace α] + [PolishSpace α] [BorelSpace α] (P : Measure α) [IsFiniteMeasure P] : + P.InnerRegularWRT (fun s ↦ IsCompact s ∧ IsClosed s) MeasurableSet := by + letI := upgradePolishSpace α + exact innerRegular_isCompact_isClosed_measurableSet_of_finite P + +end MeasureTheory diff --git a/Mathlib/MeasureTheory/Measure/Restrict.lean b/Mathlib/MeasureTheory/Measure/Restrict.lean index 71525b1a85059..cfbce99dab3ec 100644 --- a/Mathlib/MeasureTheory/Measure/Restrict.lean +++ b/Mathlib/MeasureTheory/Measure/Restrict.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro -/ import Mathlib.MeasureTheory.Measure.Comap +import Mathlib.MeasureTheory.Measure.QuasiMeasurePreserving /-! # Restricting a measure to a subset or a subtype diff --git a/Mathlib/MeasureTheory/Measure/Tilted.lean b/Mathlib/MeasureTheory/Measure/Tilted.lean index 78793a1d88732..a599a9994646b 100644 --- a/Mathlib/MeasureTheory/Measure/Tilted.lean +++ b/Mathlib/MeasureTheory/Measure/Tilted.lean @@ -119,14 +119,6 @@ lemma tilted_apply_eq_ofReal_integral [SFinite μ] (f : α → ℝ) (s : Set α) · exact ae_of_all _ (fun _ ↦ by positivity) · simp [tilted_of_not_integrable hf, integral_undef hf] -instance isFiniteMeasure_tilted : IsFiniteMeasure (μ.tilted f) := by - by_cases hf : Integrable (fun x ↦ exp (f x)) μ - · refine isFiniteMeasure_withDensity_ofReal ?_ - suffices Integrable (fun x ↦ exp (f x) / ∫ x, exp (f x) ∂μ) μ by exact this.2 - exact hf.div_const _ - · simp only [hf, not_false_eq_true, tilted_of_not_integrable] - infer_instance - lemma isProbabilityMeasure_tilted [NeZero μ] (hf : Integrable (fun x ↦ exp (f x)) μ) : IsProbabilityMeasure (μ.tilted f) := by constructor @@ -139,6 +131,16 @@ lemma isProbabilityMeasure_tilted [NeZero μ] (hf : Integrable (fun x ↦ exp (f exact integral_exp_pos hf · simp +instance isZeroOrProbabilityMeasure_tilted : IsZeroOrProbabilityMeasure (μ.tilted f) := by + rcases eq_zero_or_neZero μ with hμ | hμ + · simp only [hμ, tilted_zero_measure] + infer_instance + by_cases hf : Integrable (fun x ↦ exp (f x)) μ + · have := isProbabilityMeasure_tilted hf + infer_instance + · simp only [hf, not_false_eq_true, tilted_of_not_integrable] + infer_instance + section lintegral lemma setLIntegral_tilted' (f : α → ℝ) (g : α → ℝ≥0∞) {s : Set α} (hs : MeasurableSet s) : diff --git a/Mathlib/MeasureTheory/Order/Group/Lattice.lean b/Mathlib/MeasureTheory/Order/Group/Lattice.lean index da3f3746b9d09..6b39cae592902 100644 --- a/Mathlib/MeasureTheory/Order/Group/Lattice.lean +++ b/Mathlib/MeasureTheory/Order/Group/Lattice.lean @@ -44,6 +44,11 @@ variable [MeasurableSup₂ α] theorem measurable_mabs : Measurable (mabs : α → α) := measurable_id'.sup measurable_inv -@[to_additive (attr := measurability)] +@[to_additive (attr := measurability, fun_prop)] protected theorem Measurable.mabs (hf : Measurable f) : Measurable fun x ↦ mabs (f x) := measurable_mabs.comp hf + +@[to_additive (attr := measurability, fun_prop)] +protected theorem AEMeasurable.mabs {μ : MeasureTheory.Measure β} (hf : AEMeasurable f μ) : + AEMeasurable (fun x ↦ mabs (f x)) μ := + measurable_mabs.comp_aemeasurable hf diff --git a/Mathlib/MeasureTheory/OuterMeasure/Defs.lean b/Mathlib/MeasureTheory/OuterMeasure/Defs.lean index a776ff90e2c9c..47b1684512889 100644 --- a/Mathlib/MeasureTheory/OuterMeasure/Defs.lean +++ b/Mathlib/MeasureTheory/OuterMeasure/Defs.lean @@ -38,6 +38,8 @@ variable {α : Type*} namespace MeasureTheory +open scoped Function -- required for scoped `on` notation + /-- An outer measure is a countably subadditive monotone function that sends `∅` to `0`. -/ structure OuterMeasure (α : Type*) where /-- Outer measure function. Use automatic coercion instead. -/ diff --git a/Mathlib/MeasureTheory/PiSystem.lean b/Mathlib/MeasureTheory/PiSystem.lean index 89a46dada7446..978141449968b 100644 --- a/Mathlib/MeasureTheory/PiSystem.lean +++ b/Mathlib/MeasureTheory/PiSystem.lean @@ -482,6 +482,8 @@ end UnionInter namespace MeasurableSpace +open scoped Function -- required for scoped `on` notation + variable {α : Type*} /-! ## Dynkin systems and Π-λ theorem -/ diff --git a/Mathlib/MeasureTheory/VectorMeasure/Basic.lean b/Mathlib/MeasureTheory/VectorMeasure/Basic.lean index c316fb186a98b..a2f2b3c39872a 100644 --- a/Mathlib/MeasureTheory/VectorMeasure/Basic.lean +++ b/Mathlib/MeasureTheory/VectorMeasure/Basic.lean @@ -46,6 +46,7 @@ noncomputable section open NNReal ENNReal MeasureTheory +open scoped Function -- required for scoped `on` notation namespace MeasureTheory variable {α β : Type*} {m : MeasurableSpace α} diff --git a/Mathlib/ModelTheory/Algebra/Ring/Basic.lean b/Mathlib/ModelTheory/Algebra/Ring/Basic.lean index c2c75e666e4b2..3def47d594af0 100644 --- a/Mathlib/ModelTheory/Algebra/Ring/Basic.lean +++ b/Mathlib/ModelTheory/Algebra/Ring/Basic.lean @@ -237,7 +237,7 @@ def languageEquivEquivRingEquiv {R S : Type*} [NonAssocRing R] [NonAssocRing S] [CompatibleRing R] [CompatibleRing S] : (Language.ring.Equiv R S) ≃ (R ≃+* S) := - { toFun := fun f => + { toFun f := { f with map_add' := by intro x y @@ -245,13 +245,13 @@ def languageEquivEquivRingEquiv {R S : Type*} map_mul' := by intro x y simpa using f.map_fun mulFunc ![x, y] } - invFun := fun f => + invFun f := { f with map_fun' := fun {n} f => by cases f <;> simp map_rel' := fun {n} f => by cases f }, - left_inv := fun f => by ext; rfl - right_inv := fun f => by ext; rfl } + left_inv f := rfl + right_inv f := rfl } variable (R : Type*) [Language.ring.Structure R] diff --git a/Mathlib/ModelTheory/ElementaryMaps.lean b/Mathlib/ModelTheory/ElementaryMaps.lean index 6b3b1c267638e..b044135204fda 100644 --- a/Mathlib/ModelTheory/ElementaryMaps.lean +++ b/Mathlib/ModelTheory/ElementaryMaps.lean @@ -138,14 +138,14 @@ theorem map_constants (φ : M ↪ₑ[L] N) (c : L.Constants) : φ c = c := def toEmbedding (f : M ↪ₑ[L] N) : M ↪[L] N where toFun := f inj' := f.injective - map_fun' {_} f x := by aesop - map_rel' {_} R x := by aesop + map_fun' {_} f x := by simp + map_rel' {_} R x := by simp /-- An elementary embedding is also a first-order homomorphism. -/ def toHom (f : M ↪ₑ[L] N) : M →[L] N where toFun := f - map_fun' {_} f x := by aesop - map_rel' {_} R x := by aesop + map_fun' {_} f x := by simp + map_rel' {_} R x := by simp @[simp] theorem toEmbedding_toHom (f : M ↪ₑ[L] N) : f.toEmbedding.toHom = f.toHom := @@ -281,7 +281,7 @@ namespace Equiv /-- A first-order equivalence is also an elementary embedding. -/ def toElementaryEmbedding (f : M ≃[L] N) : M ↪ₑ[L] N where toFun := f - map_formula' n φ x := by aesop + map_formula' n φ x := by simp @[simp] theorem toElementaryEmbedding_toEmbedding (f : M ≃[L] N) : diff --git a/Mathlib/ModelTheory/LanguageMap.lean b/Mathlib/ModelTheory/LanguageMap.lean index 2daf318f51f3e..6f637d83be0fc 100644 --- a/Mathlib/ModelTheory/LanguageMap.lean +++ b/Mathlib/ModelTheory/LanguageMap.lean @@ -363,7 +363,7 @@ theorem constantsOnMap_isExpansionOn {f : α → β} {fα : α → M} {fβ : β letI := constantsOn.structure fα letI := constantsOn.structure fβ exact - ⟨fun {n} => Nat.casesOn n (fun F _x => (congr_fun h F : _)) fun n F => isEmptyElim F, fun R => + ⟨fun {n} => Nat.casesOn n (fun F _x => (congr_fun h F :)) fun n F => isEmptyElim F, fun R => isEmptyElim R⟩ end ConstantsOn diff --git a/Mathlib/ModelTheory/Substructures.lean b/Mathlib/ModelTheory/Substructures.lean index 8b6541e29a94e..6ff42480b1d52 100644 --- a/Mathlib/ModelTheory/Substructures.lean +++ b/Mathlib/ModelTheory/Substructures.lean @@ -890,8 +890,8 @@ noncomputable def substructureEquivMap (f : M ↪[L] N) (s : L.Substructure M) : (codRestrict (s.map f.toHom) (f.domRestrict s) (fun ⟨m, hm⟩ => ⟨m, hm, rfl⟩) ⟨m, hm⟩).2).2) right_inv := fun ⟨_, hn⟩ => Subtype.mk_eq_mk.2 (Classical.choose_spec hn).2 - map_fun' {n} f x := by aesop - map_rel' {n} R x := by aesop + map_fun' {n} f x := by simp + map_rel' {n} R x := by simp @[simp] theorem substructureEquivMap_apply (f : M ↪[L] N) (p : L.Substructure M) (x : p) : @@ -910,8 +910,8 @@ theorem subtype_substructureEquivMap (f : M ↪[L] N) (s : L.Substructure M) : left_inv m := f.injective (Classical.choose_spec (codRestrict f.toHom.range f f.toHom.mem_range_self m).2) right_inv := fun ⟨_, hn⟩ => Subtype.mk_eq_mk.2 (Classical.choose_spec hn) - map_fun' {n} f x := by aesop - map_rel' {n} R x := by aesop + map_fun' {n} f x := by simp + map_rel' {n} R x := by simp @[simp] theorem equivRange_apply (f : M ↪[L] N) (x : M) : (f.equivRange x : N) = f x := diff --git a/Mathlib/NumberTheory/AbelSummation.lean b/Mathlib/NumberTheory/AbelSummation.lean index 843b01ff64820..202136f5c40ad 100644 --- a/Mathlib/NumberTheory/AbelSummation.lean +++ b/Mathlib/NumberTheory/AbelSummation.lean @@ -3,7 +3,10 @@ Copyright (c) 2024 Xavier Roblot. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Xavier Roblot -/ -import Mathlib.MeasureTheory.Integral.FundThmCalculus +import Mathlib.MeasureTheory.Function.Floor +import Mathlib.MeasureTheory.Integral.Asymptotics +import Mathlib.MeasureTheory.Integral.IntegralEqImproper +import Mathlib.Topology.Order.IsLocallyClosed /-! # Abel's summation formula @@ -13,14 +16,25 @@ We prove several versions of Abel's summation formula. ## Results * `sum_mul_eq_sub_sub_integral_mul`: general statement of the formula for a sum between two -(nonnegative) reals `a` and `b`. + (nonnegative) reals `a` and `b`. * `sum_mul_eq_sub_integral_mul`: a specialized version of `sum_mul_eq_sub_sub_integral_mul` for the case `a = 0`. - * `sum_mul_eq_sub_integral_mul`: a specialized version of `sum_mul_eq_sub_integral_mul` for +* `sum_mul_eq_sub_integral_mul₀`: a specialized version of `sum_mul_eq_sub_integral_mul` for when the first coefficient of the sequence is `0`. This is useful for `ArithmeticFunction`. +Primed versions of the three results above are also stated for when the endpoints are `Nat`. + +* `tendsto_sum_mul_atTop_nhds_one_sub_integral`: limit version of `sum_mul_eq_sub_integral_mul` + when `a` tends to `∞`. + +* `tendsto_sum_mul_atTop_nhds_one_sub_integral₀`: limit version of `sum_mul_eq_sub_integral_mul₀` + when `a` tends to `∞`. + +* `summable_mul_of_bigO_atTop`: let `c : ℕ → 𝕜` and `f : ℝ → 𝕜` with `𝕜 = ℝ` or `ℂ`, prove the + summability of `n ↦ (c n) * (f n)` using Abel's formula under some `bigO` assumptions at infinity. + ## References * @@ -29,12 +43,14 @@ We prove several versions of Abel's summation formula. noncomputable section -open Finset intervalIntegral MeasureTheory IntervalIntegrable +open Finset MeasureTheory variable {𝕜 : Type*} [RCLike 𝕜] (c : ℕ → 𝕜) {f : ℝ → 𝕜} {a b : ℝ} namespace abelSummationProof +open intervalIntegral IntervalIntegrable + private theorem sumlocc (n : ℕ) : ∀ᵐ t, t ∈ Set.Icc (n : ℝ) (n + 1) → ∑ k ∈ Icc 0 ⌊t⌋₊, c k = ∑ k ∈ Icc 0 n, c k := by filter_upwards [Ico_ae_eq_Icc] with t h ht @@ -67,31 +83,40 @@ private theorem ineqofmemIco' {k : ℕ} (hk : k ∈ Ico (⌊a⌋₊ + 1) ⌊b⌋ a ≤ k ∧ k + 1 ≤ b := ineqofmemIco (by rwa [← Finset.coe_Ico]) -private theorem integrablemulsum (ha : 0 ≤ a) (hb : ⌊a⌋₊ < ⌊b⌋₊) - (hf_int : IntegrableOn (deriv f) (Set.Icc a b)) : - IntegrableOn (fun t ↦ deriv f t * (∑ k ∈ Icc 0 ⌊t⌋₊, c k)) (Set.Icc a b) := by - have h_locint {t₁ t₂ : ℝ} {n : ℕ} (h : t₁ ≤ t₂) (h₁ : n ≤ t₁) (h₂ : t₂ ≤ n + 1) - (h₃ : a ≤ t₁) (h₄ : t₂ ≤ b) : - IntervalIntegrable (fun t ↦ deriv f t * (∑ k ∈ Icc 0 ⌊t⌋₊, c k)) volume t₁ t₂ := by - rw [intervalIntegrable_iff_integrableOn_Icc_of_le h] - exact (IntegrableOn.mono_set (hf_int.mul_const _) (Set.Icc_subset_Icc h₃ h₄)).congr - <| ae_restrict_of_ae_restrict_of_subset (Set.Icc_subset_Icc h₁ h₂) - <| (ae_restrict_iff' measurableSet_Icc).mpr - (by filter_upwards [sumlocc c n] with t h ht using by rw [h ht]) - have aux1 : 0 ≤ b := (Nat.pos_of_floor_pos <| (Nat.zero_le _).trans_lt hb).le - have aux2 : ⌊a⌋₊ + 1 ≤ b := by rwa [← Nat.cast_add_one, ← Nat.le_floor_iff aux1] - have aux3 : a ≤ ⌊a⌋₊ + 1 := (Nat.lt_floor_add_one _).le - have aux4 : a ≤ ⌊b⌋₊ := le_of_lt (by rwa [← Nat.floor_lt ha]) - -- now break up into 3 subintervals - rw [← intervalIntegrable_iff_integrableOn_Icc_of_le (aux3.trans aux2)] - have I1 : IntervalIntegrable _ volume a ↑(⌊a⌋₊ + 1) := - h_locint (mod_cast aux3) (Nat.floor_le ha) (mod_cast le_rfl) le_rfl (mod_cast aux2) - have I2 : IntervalIntegrable _ volume ↑(⌊a⌋₊ + 1) ⌊b⌋₊ := - trans_iterate_Ico hb fun k hk ↦ h_locint (mod_cast k.le_succ) - le_rfl (mod_cast le_rfl) (ineqofmemIco hk).1 (mod_cast (ineqofmemIco hk).2) - have I3 : IntervalIntegrable _ volume ⌊b⌋₊ b := - h_locint (Nat.floor_le aux1) le_rfl (Nat.lt_floor_add_one _).le aux4 le_rfl - exact (I1.trans I2).trans I3 +private theorem integrablemulsum (ha : 0 ≤ a) {g : ℝ → 𝕜} (hg_int : IntegrableOn g (Set.Icc a b)) : + IntegrableOn (fun t ↦ g t * ∑ k ∈ Icc 0 ⌊t⌋₊, c k) (Set.Icc a b) := by + obtain hab | hab := le_or_gt a b + · obtain hb | hb := eq_or_lt_of_le (Nat.floor_le_floor hab) + · have : ∀ᵐ t, t ∈ Set.Icc a b → ∑ k ∈ Icc 0 ⌊a⌋₊, c k = ∑ k ∈ Icc 0 ⌊t⌋₊, c k := by + filter_upwards [sumlocc c ⌊a⌋₊] with t ht₁ ht₂ + rw [ht₁ ⟨(Nat.floor_le ha).trans ht₂.1, hb ▸ ht₂.2.trans (Nat.lt_floor_add_one b).le⟩] + rw [← ae_restrict_iff' measurableSet_Icc] at this + exact IntegrableOn.congr_fun_ae + (hg_int.mul_const _) ((Filter.EventuallyEq.refl _ g).mul this) + · have h_locint {t₁ t₂ : ℝ} {n : ℕ} (h : t₁ ≤ t₂) (h₁ : n ≤ t₁) (h₂ : t₂ ≤ n + 1) + (h₃ : a ≤ t₁) (h₄ : t₂ ≤ b) : + IntervalIntegrable (fun t ↦ g t * ∑ k ∈ Icc 0 ⌊t⌋₊, c k) volume t₁ t₂ := by + rw [intervalIntegrable_iff_integrableOn_Icc_of_le h] + exact (IntegrableOn.mono_set (hg_int.mul_const _) (Set.Icc_subset_Icc h₃ h₄)).congr + <| ae_restrict_of_ae_restrict_of_subset (Set.Icc_subset_Icc h₁ h₂) + <| (ae_restrict_iff' measurableSet_Icc).mpr + (by filter_upwards [sumlocc c n] with t h ht using by rw [h ht]) + have aux1 : 0 ≤ b := (Nat.pos_of_floor_pos <| (Nat.zero_le _).trans_lt hb).le + have aux2 : ⌊a⌋₊ + 1 ≤ b := by rwa [← Nat.cast_add_one, ← Nat.le_floor_iff aux1] + have aux3 : a ≤ ⌊a⌋₊ + 1 := (Nat.lt_floor_add_one _).le + have aux4 : a ≤ ⌊b⌋₊ := le_of_lt (by rwa [← Nat.floor_lt ha]) + -- now break up into 3 subintervals + rw [← intervalIntegrable_iff_integrableOn_Icc_of_le (aux3.trans aux2)] + have I1 : IntervalIntegrable _ volume a ↑(⌊a⌋₊ + 1) := + h_locint (mod_cast aux3) (Nat.floor_le ha) (mod_cast le_rfl) le_rfl (mod_cast aux2) + have I2 : IntervalIntegrable _ volume ↑(⌊a⌋₊ + 1) ⌊b⌋₊ := + trans_iterate_Ico hb fun k hk ↦ h_locint (mod_cast k.le_succ) + le_rfl (mod_cast le_rfl) (ineqofmemIco hk).1 (mod_cast (ineqofmemIco hk).2) + have I3 : IntervalIntegrable _ volume ⌊b⌋₊ b := + h_locint (Nat.floor_le aux1) le_rfl (Nat.lt_floor_add_one _).le aux4 le_rfl + exact (I1.trans I2).trans I3 + · rw [Set.Icc_eq_empty_of_lt hab] + exact integrableOn_empty /-- Abel's summation formula. -/ theorem _root_.sum_mul_eq_sub_sub_integral_mul (ha : 0 ≤ a) (hab : a ≤ b) @@ -99,7 +124,7 @@ theorem _root_.sum_mul_eq_sub_sub_integral_mul (ha : 0 ≤ a) (hab : a ≤ b) (hf_int : IntegrableOn (deriv f) (Set.Icc a b)) : ∑ k ∈ Ioc ⌊a⌋₊ ⌊b⌋₊, f k * c k = f b * (∑ k ∈ Icc 0 ⌊b⌋₊, c k) - f a * (∑ k ∈ Icc 0 ⌊a⌋₊, c k) - - ∫ t in Set.Ioc a b, deriv f t * (∑ k ∈ Icc 0 ⌊t⌋₊, c k) := by + ∫ t in Set.Ioc a b, deriv f t * ∑ k ∈ Icc 0 ⌊t⌋₊, c k := by rw [← integral_of_le hab] have aux1 : ⌊a⌋₊ ≤ a := Nat.floor_le ha have aux2 : b ≤ ⌊b⌋₊ + 1 := (Nat.lt_floor_add_one _).le @@ -129,45 +154,212 @@ theorem _root_.sum_mul_eq_sub_sub_integral_mul (ha : 0 ≤ a) (hab : a ≤ b) -- (Note we have 5 goals, but the 1st and 3rd are identical. TODO: find a non-hacky way of dealing -- with both at once.) · rw [intervalIntegrable_iff_integrableOn_Icc_of_le aux6] - exact (integrablemulsum c ha hb hf_int).mono_set (Set.Icc_subset_Icc_right aux5) + exact (integrablemulsum c ha hf_int).mono_set (Set.Icc_subset_Icc_right aux5) · rw [intervalIntegrable_iff_integrableOn_Icc_of_le aux5] - exact (integrablemulsum c ha hb hf_int).mono_set (Set.Icc_subset_Icc_left aux6) + exact (integrablemulsum c ha hf_int).mono_set (Set.Icc_subset_Icc_left aux6) · rw [intervalIntegrable_iff_integrableOn_Icc_of_le aux6] - exact (integrablemulsum c ha hb hf_int).mono_set (Set.Icc_subset_Icc_right aux5) + exact (integrablemulsum c ha hf_int).mono_set (Set.Icc_subset_Icc_right aux5) · rw [intervalIntegrable_iff_integrableOn_Icc_of_le aux3] - exact (integrablemulsum c ha hb hf_int).mono_set (Set.Icc_subset_Icc_right aux4) + exact (integrablemulsum c ha hf_int).mono_set (Set.Icc_subset_Icc_right aux4) · exact fun k hk ↦ (intervalIntegrable_iff_integrableOn_Icc_of_le (mod_cast k.le_succ)).mpr - <| (integrablemulsum c ha hb hf_int).mono_set + <| (integrablemulsum c ha hf_int).mono_set <| (Set.Icc_subset_Icc_iff (mod_cast k.le_succ)).mpr <| mod_cast (ineqofmemIco hk) +/-- A version of `sum_mul_eq_sub_sub_integral_mul` where the endpoints are `Nat`. -/ +theorem _root_.sum_mul_eq_sub_sub_integral_mul' {n m : ℕ} (h : n ≤ m) + (hf_diff : ∀ t ∈ Set.Icc (n : ℝ) m, DifferentiableAt ℝ f t) + (hf_int : IntegrableOn (deriv f) (Set.Icc (n : ℝ) m)) : + ∑ k ∈ Ioc n m, f k * c k = + f m * (∑ k ∈ Icc 0 m, c k) - f n * (∑ k ∈ Icc 0 n, c k) - + ∫ t in Set.Ioc (n : ℝ) m, deriv f t * ∑ k ∈ Icc 0 ⌊t⌋₊, c k := by + convert sum_mul_eq_sub_sub_integral_mul c n.cast_nonneg (Nat.cast_le.mpr h) hf_diff hf_int + all_goals rw [Nat.floor_natCast] + end abelSummationProof +section specialversions + /-- Specialized version of `sum_mul_eq_sub_sub_integral_mul` for the case `a = 0`.-/ theorem sum_mul_eq_sub_integral_mul {b : ℝ} (hb : 0 ≤ b) (hf_diff : ∀ t ∈ Set.Icc 0 b, DifferentiableAt ℝ f t) (hf_int : IntegrableOn (deriv f) (Set.Icc 0 b)) : ∑ k ∈ Icc 0 ⌊b⌋₊, f k * c k = - f b * (∑ k ∈ Icc 0 ⌊b⌋₊, c k) - ∫ t in Set.Ioc 0 b, deriv f t * (∑ k ∈ Icc 0 ⌊t⌋₊, c k) := by - nth_rewrite 1 [Finset.Icc_eq_cons_Ioc (Nat.zero_le _)] + f b * (∑ k ∈ Icc 0 ⌊b⌋₊, c k) - ∫ t in Set.Ioc 0 b, deriv f t * ∑ k ∈ Icc 0 ⌊t⌋₊, c k := by + nth_rewrite 1 [Icc_eq_cons_Ioc (Nat.zero_le _)] rw [sum_cons, ← Nat.floor_zero (α := ℝ), sum_mul_eq_sub_sub_integral_mul c le_rfl hb hf_diff hf_int, Nat.floor_zero, Nat.cast_zero, Icc_self, sum_singleton] ring +/-- A version of `sum_mul_eq_sub_integral_mul` where the endpoint is a `Nat`. -/ +theorem sum_mul_eq_sub_integral_mul' (m : ℕ) + (hf_diff : ∀ t ∈ Set.Icc (0 : ℝ) m, DifferentiableAt ℝ f t) + (hf_int : IntegrableOn (deriv f) (Set.Icc (0 : ℝ) m)) : + ∑ k ∈ Icc 0 m, f k * c k = + f m * (∑ k ∈ Icc 0 m, c k) - + ∫ t in Set.Ioc (0 : ℝ) m, deriv f t * ∑ k ∈ Icc 0 ⌊t⌋₊, c k := by + convert sum_mul_eq_sub_integral_mul c m.cast_nonneg hf_diff hf_int + all_goals rw [Nat.floor_natCast] + /-- Specialized version of `sum_mul_eq_sub_integral_mul` when the first coefficient of the sequence `c` is equal to `0`. -/ -theorem sum_mul_eq_sub_integral_mul' (hc : c 0 = 0) (b : ℝ) +theorem sum_mul_eq_sub_integral_mul₀ (hc : c 0 = 0) (b : ℝ) (hf_diff : ∀ t ∈ Set.Icc 1 b, DifferentiableAt ℝ f t) (hf_int : IntegrableOn (deriv f) (Set.Icc 1 b)) : ∑ k ∈ Icc 0 ⌊b⌋₊, f k * c k = - f b * (∑ k ∈ Icc 0 ⌊b⌋₊, c k) - ∫ t in Set.Ioc 1 b, deriv f t * (∑ k ∈ Icc 0 ⌊t⌋₊, c k) := by + f b * (∑ k ∈ Icc 0 ⌊b⌋₊, c k) - ∫ t in Set.Ioc 1 b, deriv f t * ∑ k ∈ Icc 0 ⌊t⌋₊, c k := by obtain hb | hb := le_or_gt 1 b · have : 1 ≤ ⌊b⌋₊ := (Nat.one_le_floor_iff _).mpr hb - nth_rewrite 1 [Finset.Icc_eq_cons_Ioc (by omega), sum_cons, ← Nat.Icc_succ_left, - Finset.Icc_eq_cons_Ioc (by omega), sum_cons] + nth_rewrite 1 [Icc_eq_cons_Ioc (by omega), sum_cons, ← Nat.Icc_succ_left, + Icc_eq_cons_Ioc (by omega), sum_cons] rw [Nat.succ_eq_add_one, zero_add, ← Nat.floor_one (α := ℝ), sum_mul_eq_sub_sub_integral_mul c zero_le_one hb hf_diff hf_int, Nat.floor_one, Nat.cast_one, - Finset.Icc_eq_cons_Ioc zero_le_one, sum_cons, show 1 = 0 + 1 by rfl, Nat.Ioc_succ_singleton, + Icc_eq_cons_Ioc zero_le_one, sum_cons, show 1 = 0 + 1 by rfl, Nat.Ioc_succ_singleton, zero_add, sum_singleton, hc, mul_zero, zero_add] ring · simp_rw [Nat.floor_eq_zero.mpr hb, Icc_self, sum_singleton, Nat.cast_zero, hc, mul_zero, Set.Ioc_eq_empty_of_le hb.le, Measure.restrict_empty, integral_zero_measure, sub_self] + +/-- A version of `sum_mul_eq_sub_integral_mul₀` where the endpoint is a `Nat`. -/ +theorem sum_mul_eq_sub_integral_mul₀' (hc : c 0 = 0) (m : ℕ) + (hf_diff : ∀ t ∈ Set.Icc (1 : ℝ) m, DifferentiableAt ℝ f t) + (hf_int : IntegrableOn (deriv f) (Set.Icc (1 : ℝ) m)) : + ∑ k ∈ Icc 0 m, f k * c k = + f m * (∑ k ∈ Icc 0 m, c k) - + ∫ t in Set.Ioc (1 : ℝ) m, deriv f t * ∑ k ∈ Icc 0 ⌊t⌋₊, c k := by + convert sum_mul_eq_sub_integral_mul₀ c hc m hf_diff hf_int + all_goals rw [Nat.floor_natCast] + +end specialversions + +section limit + +open Filter Topology abelSummationProof intervalIntegral + +private theorem locallyintegrablemulsum (ha : 0 ≤ a) {g : ℝ → 𝕜} (hg : IntegrableOn g (Set.Ici a)) : + LocallyIntegrableOn (fun t ↦ g t * ∑ k ∈ Icc 0 ⌊t⌋₊, c k) (Set.Ici a) := by + refine (locallyIntegrableOn_iff isLocallyClosed_Ici).mpr fun K hK₁ hK₂ ↦ ?_ + by_cases hK₃ : K.Nonempty + · have h_inf : a ≤ sInf K := (hK₁ (hK₂.sInf_mem hK₃)) + refine IntegrableOn.mono_set ?_ (Bornology.IsBounded.subset_Icc_sInf_sSup hK₂.isBounded) + refine integrablemulsum _ (ha.trans h_inf) (hg.mono_set ?_) + exact (Set.Icc_subset_Ici_iff (Real.sInf_le_sSup _ hK₂.bddBelow hK₂.bddAbove)).mpr h_inf + · rw [Set.not_nonempty_iff_eq_empty.mp hK₃] + exact integrableOn_empty + +theorem tendsto_sum_mul_atTop_nhds_one_sub_integral + (hf_diff : ∀ t ∈ Set.Ici 0, DifferentiableAt ℝ f t) + (hf_int : IntegrableOn (deriv f) (Set.Ici 0)) {l : 𝕜} + (h_lim : Tendsto (fun n : ℕ ↦ f n * ∑ k ∈ Icc 0 n, c k) atTop (𝓝 l)) + {g : ℝ → 𝕜} (hg_dom : (fun t ↦ deriv f t * ∑ k ∈ Icc 0 ⌊t⌋₊, c k) =O[atTop] g) + (hg_int : IntegrableAtFilter g atTop) : + Tendsto (fun n : ℕ ↦ ∑ k ∈ Icc 0 n, f k * c k) atTop + (𝓝 (l - ∫ t in Set.Ioi 0, deriv f t * ∑ k ∈ Icc 0 ⌊t⌋₊, c k)) := by + have h_lim' : Tendsto (fun n : ℕ ↦ ∫ t in Set.Ioc (0 : ℝ) n, deriv f t * ∑ k ∈ Icc 0 ⌊t⌋₊, c k) + atTop (𝓝 (∫ t in Set.Ioi 0, deriv f t * ∑ k ∈ Icc 0 ⌊t⌋₊, c k)) := by + refine Tendsto.congr (fun _ ↦ by rw [← integral_of_le (Nat.cast_nonneg _)]) ?_ + refine intervalIntegral_tendsto_integral_Ioi _ ?_ tendsto_natCast_atTop_atTop + exact integrableOn_Ici_iff_integrableOn_Ioi.mp + <| (locallyintegrablemulsum c le_rfl hf_int).integrableOn_of_isBigO_atTop hg_dom hg_int + refine (h_lim.sub h_lim').congr (fun _ ↦ ?_) + rw [sum_mul_eq_sub_integral_mul' _ _ (fun t ht ↦ hf_diff _ ht.1) + (hf_int.mono_set Set.Icc_subset_Ici_self)] + +theorem tendsto_sum_mul_atTop_nhds_one_sub_integral₀ (hc : c 0 = 0) + (hf_diff : ∀ t ∈ Set.Ici 1, DifferentiableAt ℝ f t) + (hf_int : IntegrableOn (deriv f) (Set.Ici 1)) {l : 𝕜} + (h_lim: Tendsto (fun n : ℕ ↦ f n * ∑ k ∈ Icc 0 n, c k) atTop (𝓝 l)) + {g : ℝ → ℝ} (hg_dom : (fun t ↦ deriv f t * ∑ k ∈ Icc 0 ⌊t⌋₊, c k) =O[atTop] g) + (hg_int : IntegrableAtFilter g atTop) : + Tendsto (fun n : ℕ ↦ ∑ k ∈ Icc 0 n, f k * c k) atTop + (𝓝 (l - ∫ t in Set.Ioi 1, deriv f t * ∑ k ∈ Icc 0 ⌊t⌋₊, c k)) := by + have h : (fun n : ℕ ↦ ∫ (x : ℝ) in (1 : ℝ)..n, deriv f x * ∑ k ∈ Icc 0 ⌊x⌋₊, c k) =ᶠ[atTop] + (fun n : ℕ ↦ ∫ (t : ℝ) in Set.Ioc 1 ↑n, deriv f t * ∑ k ∈ Icc 0 ⌊t⌋₊, c k) := by + filter_upwards [eventually_ge_atTop 1] with _ h + rw [← integral_of_le (Nat.one_le_cast.mpr h)] + have h_lim' : Tendsto (fun n : ℕ ↦ ∫ t in Set.Ioc (1 : ℝ) n, deriv f t * ∑ k ∈ Icc 0 ⌊t⌋₊, c k) + atTop (𝓝 (∫ t in Set.Ioi 1, deriv f t * ∑ k ∈ Icc 0 ⌊t⌋₊, c k)) := by + refine Tendsto.congr' h (intervalIntegral_tendsto_integral_Ioi _ ?_ tendsto_natCast_atTop_atTop) + exact integrableOn_Ici_iff_integrableOn_Ioi.mp + <| (locallyintegrablemulsum c zero_le_one hf_int).integrableOn_of_isBigO_atTop hg_dom hg_int + refine (h_lim.sub h_lim').congr (fun _ ↦ ?_) + rw [sum_mul_eq_sub_integral_mul₀' _ hc _ (fun t ht ↦ hf_diff _ ht.1) + (hf_int.mono_set Set.Icc_subset_Ici_self)] + +end limit + +section summable + +open Filter abelSummationProof + +private theorem summable_mul_of_bigO_atTop_aux (m : ℕ) + (h_bdd : (fun n : ℕ ↦ ‖f n‖ * ∑ k ∈ Icc 0 n, ‖c k‖) =O[atTop] fun _ ↦ (1 : ℝ)) + (hf_int : IntegrableOn (deriv (fun t ↦ ‖f t‖)) (Set.Ici (m : ℝ))) + (hf : ∀ n : ℕ, ∑ k ∈ Icc 0 n, ‖f k‖ * ‖c k‖ = + ‖f n‖ * ∑ k ∈ Icc 0 n, ‖c k‖ - + ∫ (t : ℝ) in Set.Ioc ↑m ↑n, deriv (fun t ↦ ‖f t‖) t * ∑ k ∈ Icc 0 ⌊t⌋₊, ‖c k‖) + {g : ℝ → ℝ} + (hg₁ : (fun t ↦ deriv (fun t ↦ ‖f t‖) t * ∑ k ∈ Icc 0 ⌊t⌋₊, ‖c k‖) =O[atTop] g) + (hg₂ : IntegrableAtFilter g atTop) : + Summable (fun n : ℕ ↦ f n * c n) := by + obtain ⟨C₁, hC₁⟩ := Asymptotics.isBigO_one_nat_atTop_iff.mp h_bdd + let C₂ := ∫ t in Set.Ioi (m : ℝ), ‖deriv (fun t ↦ ‖f t‖) t * ∑ k ∈ Icc 0 ⌊t⌋₊, ‖c k‖‖ + refine summable_of_sum_range_norm_le (c := max (C₁ + C₂) 1) fun n ↦ ?_ + cases n with + | zero => simp only [range_zero, norm_mul, sum_empty, le_sup_iff, zero_le_one, or_true] + | succ n => + have h_mes : Measurable fun t ↦ deriv (fun t ↦ ‖f t‖) t * ∑ k ∈ Icc 0 ⌊t⌋₊, ‖c k‖ := + (measurable_deriv _).mul <| Measurable.comp' (g := fun n : ℕ ↦ ∑ k ∈ Icc 0 n, ‖c k‖) + (fun _ _ ↦ trivial) Nat.measurable_floor + rw [Nat.range_eq_Icc_zero_sub_one _ n.add_one_ne_zero, add_tsub_cancel_right] + calc + _ = ∑ k ∈ Icc 0 n, ‖f k‖ * ‖c k‖ := by simp_rw [norm_mul] + _ = ‖f n‖ * ∑ k ∈ Icc 0 n, ‖c k‖ - + ∫ t in Set.Ioc ↑m ↑n, deriv (fun t ↦ ‖f t‖) t * ∑ k ∈ Icc 0 ⌊t⌋₊, ‖c k‖ := ?_ + _ ≤ C₁ - ∫ t in Set.Ioc ↑m ↑n, deriv (fun t ↦ ‖f t‖) t * ∑ k ∈ Icc 0 ⌊t⌋₊, ‖c k‖ := ?_ + _ ≤ C₁ + ∫ t in Set.Ioc ↑m ↑n, ‖deriv (fun t ↦ ‖f t‖) t * ∑ k ∈ Icc 0 ⌊t⌋₊, ‖c k‖‖ := ?_ + _ ≤ C₁ + C₂ := ?_ + _ ≤ max (C₁ + C₂) 1 := le_max_left _ _ + · exact hf _ + · refine tsub_le_tsub_right (le_of_eq_of_le (Real.norm_of_nonneg ?_).symm (hC₁ n)) _ + exact mul_nonneg (norm_nonneg _) (sum_nonneg fun _ _ ↦ norm_nonneg _) + · exact add_le_add_left + (le_trans (neg_le_abs _) (Real.norm_eq_abs _ ▸ norm_integral_le_integral_norm _)) _ + · refine add_le_add_left (setIntegral_mono_set ?_ ?_ Set.Ioc_subset_Ioi_self.eventuallyLE) C₁ + · refine integrableOn_Ici_iff_integrableOn_Ioi.mp <| + (integrable_norm_iff h_mes.aestronglyMeasurable).mpr <| + (locallyintegrablemulsum _ m.cast_nonneg hf_int).integrableOn_of_isBigO_atTop hg₁ hg₂ + · filter_upwards with t using norm_nonneg _ + +theorem summable_mul_of_bigO_atTop + (hf_diff : ∀ t ∈ Set.Ici 0, DifferentiableAt ℝ (fun x ↦ ‖f x‖) t) + (hf_int : IntegrableOn (deriv (fun t ↦ ‖f t‖)) (Set.Ici 0)) + (h_bdd : (fun n : ℕ ↦ ‖f n‖ * ∑ k ∈ Icc 0 n, ‖c k‖) =O[atTop] fun _ ↦ (1 : ℝ)) + {g : ℝ → ℝ} (hg₁ : (fun t ↦ deriv (fun t ↦ ‖f t‖) t * ∑ k ∈ Icc 0 ⌊t⌋₊, ‖c k‖) =O[atTop] g) + (hg₂ : IntegrableAtFilter g atTop) : + Summable (fun n : ℕ ↦ f n * c n) := by + refine summable_mul_of_bigO_atTop_aux c 0 h_bdd (by rwa [Nat.cast_zero]) (fun n ↦ ?_) hg₁ hg₂ + exact_mod_cast sum_mul_eq_sub_integral_mul' _ _ (fun _ ht ↦ hf_diff _ ht.1) + (hf_int.mono_set Set.Icc_subset_Ici_self) + +/-- A version of `summable_mul_of_bigO_atTop` that can be useful to avoid difficulties near zero. -/ +theorem summable_mul_of_bigO_atTop' + (hf_diff : ∀ t ∈ Set.Ici 1, DifferentiableAt ℝ (fun x ↦ ‖f x‖) t) + (hf_int : IntegrableOn (deriv (fun t ↦ ‖f t‖)) (Set.Ici 1)) + (h_bdd : (fun n : ℕ ↦ ‖f n‖ * ∑ k ∈ Icc 1 n, ‖c k‖) =O[atTop] fun _ ↦ (1 : ℝ)) + {g : ℝ → ℝ} (hg₁ : (fun t ↦ deriv (fun t ↦ ‖f t‖) t * ∑ k ∈ Icc 1 ⌊t⌋₊, ‖c k‖) =O[atTop] g) + (hg₂ : IntegrableAtFilter g atTop) : + Summable (fun n : ℕ ↦ f n * c n) := by + have h : ∀ n, ∑ k ∈ Icc 1 n, ‖c k‖ = ∑ k ∈ Icc 0 n, ‖(fun n ↦ if n = 0 then 0 else c n) k‖ := by + intro n + rw [Icc_eq_cons_Ioc n.zero_le, sum_cons, ← Nat.Icc_succ_left, Nat.succ_eq_add_one, zero_add] + simp_rw [if_pos, norm_zero, zero_add] + exact Finset.sum_congr rfl fun _ h ↦ by rw [if_neg (zero_lt_one.trans_le (mem_Icc.mp h).1).ne'] + simp_rw [h] at h_bdd hg₁ + refine Summable.congr_atTop (summable_mul_of_bigO_atTop_aux (fun n ↦ if n = 0 then 0 else c n) 1 + h_bdd (by rwa [Nat.cast_one]) (fun n ↦ ?_) hg₁ hg₂) ?_ + · exact_mod_cast sum_mul_eq_sub_integral_mul₀' _ (by simp only [reduceIte, norm_zero]) n + (fun _ ht ↦ hf_diff _ ht.1) (hf_int.mono_set Set.Icc_subset_Ici_self) + · filter_upwards [eventually_ne_atTop 0] with k hk + simp_rw [if_neg hk] + +end summable diff --git a/Mathlib/NumberTheory/ArithmeticFunction.lean b/Mathlib/NumberTheory/ArithmeticFunction.lean index 900057426c931..e8781f75262c5 100644 --- a/Mathlib/NumberTheory/ArithmeticFunction.lean +++ b/Mathlib/NumberTheory/ArithmeticFunction.lean @@ -568,6 +568,7 @@ theorem map_mul_of_coprime {f : ArithmeticFunction R} (hf : f.IsMultiplicative) end MonoidWithZero +open scoped Function in -- required for scoped `on` notation theorem map_prod {ι : Type*} [CommMonoidWithZero R] (g : ι → ℕ) {f : ArithmeticFunction R} (hf : f.IsMultiplicative) (s : Finset ι) (hs : (s : Set ι).Pairwise (Coprime on g)) : f (∏ i ∈ s, g i) = ∏ i ∈ s, f (g i) := by diff --git a/Mathlib/NumberTheory/Cyclotomic/Basic.lean b/Mathlib/NumberTheory/Cyclotomic/Basic.lean index ec3877fb1cd1d..c9e7e4599210e 100644 --- a/Mathlib/NumberTheory/Cyclotomic/Basic.lean +++ b/Mathlib/NumberTheory/Cyclotomic/Basic.lean @@ -548,7 +548,7 @@ instance CyclotomicField.noZeroSMulDivisors [IsFractionRing A K] : rw [IsScalarTower.algebraMap_eq A K (CyclotomicField n K)] exact (Function.Injective.comp (NoZeroSMulDivisors.algebraMap_injective K (CyclotomicField n K)) - (IsFractionRing.injective A K) : _) + (IsFractionRing.injective A K) :) /-- If `A` is a domain with fraction field `K` and `n : ℕ+`, we define `CyclotomicRing n A K` as the `A`-subalgebra of `CyclotomicField n K` generated by the roots of `X ^ n - 1`. If `n` diff --git a/Mathlib/NumberTheory/FLT/MasonStothers.lean b/Mathlib/NumberTheory/FLT/MasonStothers.lean index eee4ff81ac138..bd5d8a2c3a537 100644 --- a/Mathlib/NumberTheory/FLT/MasonStothers.lean +++ b/Mathlib/NumberTheory/FLT/MasonStothers.lean @@ -16,9 +16,6 @@ Proof is based on this online note by Franz Lemmermeyer http://www.fen.bilkent.e which is essentially based on Noah Snyder's paper "An Alternative Proof of Mason's Theorem", but slightly different. -## TODO - -Prove polynomial FLT using Mason-Stothers theorem. -/ open Polynomial UniqueFactorizationMonoid UniqueFactorizationDomain EuclideanDomain @@ -51,13 +48,25 @@ private theorem abc_subcall {a b c w : k[X]} {hw : w ≠ 0} (wab : w = wronskian exact Nat.add_lt_add_right abc_dr_ndeg_lt _ /-- **Polynomial ABC theorem.** -/ -theorem Polynomial.abc {a b c : k[X]} (ha : a ≠ 0) (hb : b ≠ 0) (hc : c ≠ 0) (hab : IsCoprime a b) - (hbc : IsCoprime b c) (hca : IsCoprime c a) (hsum : a + b + c = 0) : +protected theorem Polynomial.abc + {a b c : k[X]} (ha : a ≠ 0) (hb : b ≠ 0) (hc : c ≠ 0) + (hab : IsCoprime a b) (hsum : a + b + c = 0) : ( natDegree a + 1 ≤ (radical (a * b * c)).natDegree ∧ natDegree b + 1 ≤ (radical (a * b * c)).natDegree ∧ natDegree c + 1 ≤ (radical (a * b * c)).natDegree ) ∨ derivative a = 0 ∧ derivative b = 0 ∧ derivative c = 0 := by set w := wronskian a b with wab + have hbc : IsCoprime b c := by + rw [add_eq_zero_iff_neg_eq] at hsum + rw [← hsum, IsCoprime.neg_right_iff] + convert IsCoprime.add_mul_left_right hab.symm 1 + rw [mul_one] + have hsum' : b + c + a = 0 := by rwa [add_rotate] at hsum + have hca : IsCoprime c a := by + rw [add_eq_zero_iff_neg_eq] at hsum' + rw [← hsum', IsCoprime.neg_right_iff] + convert IsCoprime.add_mul_left_right hbc.symm 1 + rw [mul_one] have wbc : w = wronskian b c := wronskian_eq_of_sum_zero hsum have wca : w = wronskian c a := by rw [add_rotate] at hsum diff --git a/Mathlib/NumberTheory/FLT/Polynomial.lean b/Mathlib/NumberTheory/FLT/Polynomial.lean new file mode 100644 index 0000000000000..e6beb32165d5a --- /dev/null +++ b/Mathlib/NumberTheory/FLT/Polynomial.lean @@ -0,0 +1,275 @@ +/- +Copyright (c) 2024 Jineon Baek and Seewoo Lee. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jineon Baek, Seewoo Lee +-/ +import Mathlib.Algebra.Polynomial.Expand +import Mathlib.Algebra.GroupWithZero.Defs +import Mathlib.NumberTheory.FLT.Basic +import Mathlib.NumberTheory.FLT.MasonStothers +import Mathlib.RingTheory.Polynomial.Content + +/-! +# Fermat's Last Theorem for polynomials over a field + +This file states and proves the Fermat's Last Theorem for polynomials over a field. +For `n ≥ 3` not divisible by the characteristic of the coefficient field `k` and (pairwise) nonzero +coprime polynomials `a, b, c` (over a field) with `a ^ n + b ^ n = c ^ n`, +all polynomials must be constants. + +More generally, we can prove non-solvability of the Fermat-Catalan equation: there are no +non-constant polynomial solutions to the equation `u * a ^ p + v * b ^ q + w * c ^ r = 0`, where +`p, q, r ≥ 3` with `p * q + q * r + r * p ≤ p * q * r` , `p, q, r` not divisible by `char k`, +and `u, v, w` are nonzero elements in `k`. +FLT is the special case where `p = q = r = n`, `u = v = 1`, and `w = -1`. + +The proof uses the Mason-Stothers theorem (Polynomial ABC theorem) and infinite descent +(in the characteristic p case). +-/ + +open Polynomial UniqueFactorizationMonoid UniqueFactorizationDomain + +variable {k R : Type*} [Field k] [CommRing R] [IsDomain R] [NormalizationMonoid R] + [UniqueFactorizationMonoid R] + +private lemma Ne.isUnit_C {u : k} (hu : u ≠ 0) : IsUnit (C u) := + Polynomial.isUnit_C.mpr hu.isUnit + +-- auxiliary lemma that 'rotates' coprimality +private lemma rot_coprime + {p q r : ℕ} {a b c : k[X]} {u v w : k} + {hp : 0 < p} {hq : 0 < q} {hr : 0 < r} + {hu : u ≠ 0} {hv : v ≠ 0} {hw : w ≠ 0} + (heq : C u * a ^ p + C v * b ^ q + C w * c ^ r = 0) (hab : IsCoprime a b) : IsCoprime b c := by + have hCu : IsUnit (C u) := Ne.isUnit_C hu + have hCv : IsUnit (C v) := Ne.isUnit_C hv + have hCw : IsUnit (C w) := Ne.isUnit_C hw + rw [← IsCoprime.pow_iff hp hq, ← isCoprime_mul_units_left hCu hCv] at hab + rw [add_eq_zero_iff_neg_eq] at heq + rw [← IsCoprime.pow_iff hq hr, ← isCoprime_mul_units_left hCv hCw, + ← heq, IsCoprime.neg_right_iff] + convert IsCoprime.add_mul_left_right hab.symm 1 using 2 + rw [mul_one] + +private lemma ineq_pqr_contradiction {p q r a b c : ℕ} + (hp : 0 < p) (hq : 0 < q) (hr : 0 < r) + (hineq : q * r + r * p + p * q ≤ p * q * r) + (hpa : p * a < a + b + c) + (hqb : q * b < a + b + c) + (hrc : r * c < a + b + c) : False := by + suffices h : p * q * r * (a + b + c) + 1 ≤ p * q * r * (a + b + c) by simp at h + calc + _ = (p * a) * (q * r) + (q * b) * (r * p) + (r * c) * (p * q) + 1 := by ring + _ ≤ (a + b + c) * (q * r) + (a + b + c) * (r * p) + (a + b + c) * (p * q) := by + rw [Nat.succ_le] + gcongr + _ = (q * r + r * p + p * q) * (a + b + c) := by ring + _ ≤ _ := by gcongr + +private theorem Polynomial.flt_catalan_deriv [DecidableEq k] + {p q r : ℕ} (hp : 0 < p) (hq : 0 < q) (hr : 0 < r) + (hineq : q * r + r * p + p * q ≤ p * q * r) + (chp : (p : k) ≠ 0) (chq : (q : k) ≠ 0) (chr : (r : k) ≠ 0) + {a b c : k[X]} (ha : a ≠ 0) (hb : b ≠ 0) (hc : c ≠ 0) + (hab : IsCoprime a b) {u v w : k} (hu : u ≠ 0) (hv : v ≠ 0) (hw : w ≠ 0) + (heq : C u * a ^ p + C v * b ^ q + C w * c ^ r = 0) : + derivative a = 0 ∧ derivative b = 0 ∧ derivative c = 0 := by + have hbc : IsCoprime b c := by apply rot_coprime heq <;> assumption + have hca : IsCoprime c a := by + rw [add_rotate] at heq; apply rot_coprime heq <;> assumption + have hCu := C_ne_zero.mpr hu + have hCv := C_ne_zero.mpr hv + have hCw := C_ne_zero.mpr hw + have hap := pow_ne_zero p ha + have hbq := pow_ne_zero q hb + have hcr := pow_ne_zero r hc + have habp : IsCoprime (C u * a ^ p) (C v * b ^ q) := by + rw [isCoprime_mul_units_left hu.isUnit_C hv.isUnit_C]; exact hab.pow + have hbcp : IsCoprime (C v * b ^ q) (C w * c ^ r) := by + rw [isCoprime_mul_units_left hv.isUnit_C hw.isUnit_C]; exact hbc.pow + have hcap : IsCoprime (C w * c ^ r) (C u * a ^ p) := by + rw [isCoprime_mul_units_left hw.isUnit_C hu.isUnit_C]; exact hca.pow + have habcp := hcap.symm.mul_left hbcp + + -- Use Mason-Stothers theorem + rcases Polynomial.abc + (mul_ne_zero hCu hap) (mul_ne_zero hCv hbq) (mul_ne_zero hCw hcr) + habp heq with nd_lt | dr0 + · simp_rw [radical_mul habcp, radical_mul habp, + radical_mul_of_isUnit_left hu.isUnit_C, + radical_mul_of_isUnit_left hv.isUnit_C, + radical_mul_of_isUnit_left hw.isUnit_C, + radical_pow a hp, radical_pow b hq, radical_pow c hr, + natDegree_mul hCu hap, + natDegree_mul hCv hbq, + natDegree_mul hCw hcr, + natDegree_C, natDegree_pow, zero_add, + ← radical_mul hab, + ← radical_mul (hca.symm.mul_left hbc)] at nd_lt + + obtain ⟨hpa', hqb', hrc'⟩ := nd_lt + have hpa := hpa'.trans natDegree_radical_le + have hqb := hqb'.trans natDegree_radical_le + have hrc := hrc'.trans natDegree_radical_le + rw [natDegree_mul (mul_ne_zero ha hb) hc, + natDegree_mul ha hb, Nat.add_one_le_iff] at hpa hqb hrc + exfalso + exact (ineq_pqr_contradiction hp hq hr hineq hpa hqb hrc) + · rw [derivative_C_mul, derivative_C_mul, derivative_C_mul, + mul_eq_zero_iff_left (C_ne_zero.mpr hu), + mul_eq_zero_iff_left (C_ne_zero.mpr hv), + mul_eq_zero_iff_left (C_ne_zero.mpr hw), + derivative_pow_eq_zero chp, + derivative_pow_eq_zero chq, + derivative_pow_eq_zero chr] at dr0 + exact dr0 + +-- helper lemma that gives a baggage of small facts on `contract (ringChar k) a` +private lemma find_contract {a : k[X]} + (ha : a ≠ 0) (hda : derivative a = 0) (chn0 : ringChar k ≠ 0) : + ∃ ca, ca ≠ 0 ∧ + a = expand k (ringChar k) ca ∧ + a.natDegree = ca.natDegree * ringChar k := by + have heq := (expand_contract (ringChar k) hda chn0).symm + refine ⟨contract (ringChar k) a, ?_, heq, ?_⟩ + · intro h + rw [h, map_zero] at heq + exact ha heq + · rw [← natDegree_expand, ← heq] + +variable [DecidableEq k] + +private theorem Polynomial.flt_catalan_aux + {p q r : ℕ} {a b c : k[X]} {u v w : k} + (heq : C u * a ^ p + C v * b ^ q + C w * c ^ r = 0) + (hp : 0 < p) (hq : 0 < q) (hr : 0 < r) + (hineq : q * r + r * p + p * q ≤ p * q * r) + (chp : (p : k) ≠ 0) (chq : (q : k) ≠ 0) (chr : (r : k) ≠ 0) + (ha : a ≠ 0) (hb : b ≠ 0) (hc : c ≠ 0) (hab : IsCoprime a b) + (hu : u ≠ 0) (hv : v ≠ 0) (hw : w ≠ 0) : + a.natDegree = 0 := by + cases' eq_or_ne (ringChar k) 0 with ch0 chn0 + -- characteristic zero + · obtain ⟨da, -, -⟩ := flt_catalan_deriv + hp hq hr hineq chp chq chr ha hb hc hab hu hv hw heq + have czk : CharZero k := by + apply charZero_of_inj_zero + intro n + rw [ringChar.spec, ch0] + exact zero_dvd_iff.mp + rw [eq_C_of_derivative_eq_zero da] + exact natDegree_C _ + -- characteristic p > 0 + · generalize eq_d : a.natDegree = d + -- set up infinite descent + -- strong induct on `d := a.natDegree` + induction d + using Nat.case_strong_induction_on + generalizing a b c ha hb hc hab heq with + | hz => rfl + | hi d ih_d => -- have derivatives of `a, b, c` zero + obtain ⟨ad, bd, cd⟩ := flt_catalan_deriv + hp hq hr hineq chp chq chr ha hb hc hab hu hv hw heq + -- find contracts `ca, cb, cc` so that `a(k) = ca(k^ch)` + obtain ⟨ca, ca_nz, eq_a, eq_deg_a⟩ := find_contract ha ad chn0 + obtain ⟨cb, cb_nz, eq_b, eq_deg_b⟩ := find_contract hb bd chn0 + obtain ⟨cc, cc_nz, eq_c, eq_deg_c⟩ := find_contract hc cd chn0 + set ch := ringChar k + suffices hca : ca.natDegree = 0 by + rw [← eq_d, eq_deg_a, hca, zero_mul] + by_contra hnca; apply hnca + apply ih_d _ _ _ ca_nz cb_nz cc_nz <;> clear ih_d <;> try rfl + · apply (isCoprime_expand chn0).mp + rwa [← eq_a, ← eq_b] + · have _ : ch ≠ 1 := CharP.ringChar_ne_one + have hch2 : 2 ≤ ch := by omega + rw [← add_le_add_iff_right 1, ← eq_d, eq_deg_a] + refine le_trans ?_ (Nat.mul_le_mul_left _ hch2) + omega + · rw [eq_a, eq_b, eq_c, ← expand_C ch u, ← expand_C ch v, ← expand_C ch w] at heq + simp_rw [← map_pow, ← map_mul, ← map_add] at heq + rwa [Polynomial.expand_eq_zero (zero_lt_iff.mpr chn0)] at heq + +/-- Nonsolvability of the Fermat-Catalan equation. -/ +theorem Polynomial.flt_catalan + {p q r : ℕ} (hp : 0 < p) (hq : 0 < q) (hr : 0 < r) + (hineq : q * r + r * p + p * q ≤ p * q * r) + (chp : (p : k) ≠ 0) (chq : (q : k) ≠ 0) (chr : (r : k) ≠ 0) + {a b c : k[X]} (ha : a ≠ 0) (hb : b ≠ 0) (hc : c ≠ 0) (hab : IsCoprime a b) + {u v w : k} (hu : u ≠ 0) (hv : v ≠ 0) (hw : w ≠ 0) + (heq : C u * a ^ p + C v * b ^ q + C w * c ^ r = 0) : + a.natDegree = 0 ∧ b.natDegree = 0 ∧ c.natDegree = 0 := by + -- WLOG argument: essentially three times flt_catalan_aux + have hbc : IsCoprime b c := by + apply rot_coprime heq hab <;> assumption + have heq' : C v * b ^ q + C w * c ^ r + C u * a ^ p = 0 := by + rwa [add_rotate] at heq + have hca : IsCoprime c a := by + apply rot_coprime heq' hbc <;> assumption + refine ⟨?_, ?_, ?_⟩ + · apply flt_catalan_aux heq <;> assumption + · rw [add_rotate] at heq hineq + rw [mul_rotate] at hineq + apply flt_catalan_aux heq <;> assumption + · rw [← add_rotate] at heq hineq + rw [← mul_rotate] at hineq + apply flt_catalan_aux heq <;> assumption + +theorem Polynomial.flt + {n : ℕ} (hn : 3 ≤ n) (chn : (n : k) ≠ 0) + {a b c : k[X]} (ha : a ≠ 0) (hb : b ≠ 0) (hc : c ≠ 0) + (hab : IsCoprime a b) (heq : a ^ n + b ^ n = c ^ n) : + a.natDegree = 0 ∧ b.natDegree = 0 ∧ c.natDegree = 0 := by + have hn' : 0 < n := by linarith + rw [← sub_eq_zero, ← one_mul (a ^ n), ← one_mul (b ^ n), ← one_mul (c ^ n), sub_eq_add_neg, + ← neg_mul] at heq + have hone : (1 : k[X]) = C 1 := by rfl + have hneg_one : (-1 : k[X]) = C (-1) := by simp only [map_neg, map_one] + simp_rw [hneg_one, hone] at heq + apply flt_catalan hn' hn' hn' _ + chn chn chn ha hb hc hab one_ne_zero one_ne_zero (neg_ne_zero.mpr one_ne_zero) heq + have eq_lhs : n * n + n * n + n * n = 3 * n * n := by ring + rw [eq_lhs, mul_assoc, mul_assoc] + exact Nat.mul_le_mul_right (n * n) hn + +theorem fermatLastTheoremWith'_polynomial {n : ℕ} (hn : 3 ≤ n) (chn : (n : k) ≠ 0) : + FermatLastTheoremWith' k[X] n := by + rw [FermatLastTheoremWith'] + intros a b c ha hb hc heq + obtain ⟨a', eq_a⟩ := gcd_dvd_left a b + obtain ⟨b', eq_b⟩ := gcd_dvd_right a b + set d := gcd a b + have hd : d ≠ 0 := gcd_ne_zero_of_left ha + rw [eq_a, eq_b, mul_pow, mul_pow, ← mul_add] at heq + have hdc : d ∣ c := by + have hn : 0 < n := by omega + have hdncn : d ^ n ∣ c ^ n := ⟨_, heq.symm⟩ + + rw [dvd_iff_normalizedFactors_le_normalizedFactors hd hc] + rw [dvd_iff_normalizedFactors_le_normalizedFactors + (pow_ne_zero n hd) (pow_ne_zero n hc), + normalizedFactors_pow, normalizedFactors_pow] at hdncn + simp_rw [Multiset.le_iff_count, Multiset.count_nsmul, + mul_le_mul_left hn] at hdncn ⊢ + exact hdncn + obtain ⟨c', eq_c⟩ := hdc + rw [eq_a, mul_ne_zero_iff] at ha + rw [eq_b, mul_ne_zero_iff] at hb + rw [eq_c, mul_ne_zero_iff] at hc + rw [mul_comm] at eq_a eq_b eq_c + refine ⟨d, a', b', c', ⟨eq_a, eq_b, eq_c⟩, ?_⟩ + rw [eq_c, mul_pow, mul_comm, mul_left_inj' (pow_ne_zero n hd)] at heq + suffices goal : a'.natDegree = 0 ∧ b'.natDegree = 0 ∧ c'.natDegree = 0 by + simp [natDegree_eq_zero] at goal + obtain ⟨⟨ca', ha'⟩, ⟨cb', hb'⟩, ⟨cc', hc'⟩⟩ := goal + rw [← ha', ← hb', ← hc'] + rw [← ha', C_ne_zero] at ha + rw [← hb', C_ne_zero] at hb + rw [← hc', C_ne_zero] at hc + exact ⟨ha.right.isUnit_C, hb.right.isUnit_C, hc.right.isUnit_C⟩ + apply flt hn chn ha.right hb.right hc.right _ heq + convert isCoprime_div_gcd_div_gcd _ + · exact EuclideanDomain.eq_div_of_mul_eq_left ha.left eq_a.symm + · exact EuclideanDomain.eq_div_of_mul_eq_left ha.left eq_b.symm + · rw [eq_b] + exact mul_ne_zero hb.right hb.left diff --git a/Mathlib/NumberTheory/LSeries/Dirichlet.lean b/Mathlib/NumberTheory/LSeries/Dirichlet.lean index 32e114a35c3d6..96724285e05ad 100644 --- a/Mathlib/NumberTheory/LSeries/Dirichlet.lean +++ b/Mathlib/NumberTheory/LSeries/Dirichlet.lean @@ -264,9 +264,6 @@ lemma LSeries_zeta_eq : L ↗ζ = L 1 := by theorem LSeriesSummable_zeta_iff {s : ℂ} : LSeriesSummable (ζ ·) s ↔ 1 < s.re := (LSeriesSummable_congr s const_one_eq_zeta).symm.trans <| LSeriesSummable_one_iff -@[deprecated (since := "2024-03-29")] -alias zeta_LSeriesSummable_iff_one_lt_re := LSeriesSummable_zeta_iff - /-- The abscissa of (absolute) convergence of the arithmetic function `ζ` is `1`. -/ lemma abscissaOfAbsConv_zeta : abscissaOfAbsConv ↗ζ = 1 := by rw [abscissaOfAbsConv_congr (g := 1) fun hn ↦ by simp [hn], abscissaOfAbsConv_one] diff --git a/Mathlib/NumberTheory/LucasLehmer.lean b/Mathlib/NumberTheory/LucasLehmer.lean index 7990d0b3b0785..01356b10609e7 100644 --- a/Mathlib/NumberTheory/LucasLehmer.lean +++ b/Mathlib/NumberTheory/LucasLehmer.lean @@ -270,14 +270,12 @@ instance : NatCast (X q) where @[simp] theorem snd_natCast (n : ℕ) : (n : X q).snd = (0 : ZMod q) := rfl --- See note [no_index around OfNat.ofNat] @[simp] theorem ofNat_fst (n : ℕ) [n.AtLeastTwo] : - (no_index (OfNat.ofNat n) : X q).fst = OfNat.ofNat n := + (ofNat(n) : X q).fst = OfNat.ofNat n := rfl --- See note [no_index around OfNat.ofNat] @[simp] theorem ofNat_snd (n : ℕ) [n.AtLeastTwo] : - (no_index (OfNat.ofNat n) : X q).snd = 0 := + (ofNat(n) : X q).snd = 0 := rfl instance : AddGroupWithOne (X q) := diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Defs.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Defs.lean index 890e52757badd..75d79867aa03b 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Defs.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Defs.lean @@ -38,6 +38,7 @@ section gammaSet_def /-- The set of pairs of coprime integers congruent to `a` mod `N`. -/ def gammaSet := {v : Fin 2 → ℤ | (↑) ∘ v = a ∧ IsCoprime (v 0) (v 1)} +open scoped Function in -- required for scoped `on` notation lemma pairwise_disjoint_gammaSet : Pairwise (Disjoint on gammaSet N) := by refine fun u v huv ↦ ?_ contrapose! huv diff --git a/Mathlib/NumberTheory/NumberField/AdeleRing.lean b/Mathlib/NumberTheory/NumberField/AdeleRing.lean index 1ee25ce15fd5f..927332bf8bdeb 100644 --- a/Mathlib/NumberTheory/NumberField/AdeleRing.lean +++ b/Mathlib/NumberTheory/NumberField/AdeleRing.lean @@ -41,8 +41,6 @@ open InfinitePlace AbsoluteValue.Completion InfinitePlace.Completion DedekindDom open scoped Classical -variable (K : Type*) [Field K] - /-! ## The infinite adele ring The infinite adele ring is the finite product of completions of a number field over its @@ -51,10 +49,12 @@ infinite places. See `NumberField.InfinitePlace` for the definition of an infini -/ /-- The infinite adele ring of a number field. -/ -def InfiniteAdeleRing := (v : InfinitePlace K) → v.Completion +def InfiniteAdeleRing (K : Type*) [Field K] := (v : InfinitePlace K) → v.Completion namespace InfiniteAdeleRing +variable (K : Type*) [Field K] + instance : CommRing (InfiniteAdeleRing K) := Pi.commRing instance : Inhabited (InfiniteAdeleRing K) := ⟨0⟩ @@ -115,38 +115,47 @@ theorem mixedEmbedding_eq_algebraMap_comp {x : K} : end InfiniteAdeleRing -variable [NumberField K] - /-! ## The adele ring -/ -/-- The adele ring of a number field. -/ -def AdeleRing := InfiniteAdeleRing K × FiniteAdeleRing (𝓞 K) K +/-- `AdeleRing (𝓞 K) K` is the adele ring of a number field `K`. + +More generally `AdeleRing R K` can be used if `K` is the field of fractions +of the Dedekind domain `R`. This enables use of rings like `AdeleRing ℤ ℚ`, which +in practice are easier to work with than `AdeleRing (𝓞 ℚ) ℚ`. + +Note that this definition does not give the correct answer in the function field case. +-/ +def AdeleRing (R K : Type*) [CommRing R] [IsDedekindDomain R] [Field K] + [Algebra R K] [IsFractionRing R K] := InfiniteAdeleRing K × FiniteAdeleRing R K namespace AdeleRing -instance : CommRing (AdeleRing K) := Prod.instCommRing +variable (R K : Type*) [CommRing R] [IsDedekindDomain R] [Field K] + [Algebra R K] [IsFractionRing R K] + +instance : CommRing (AdeleRing R K) := Prod.instCommRing -instance : Inhabited (AdeleRing K) := ⟨0⟩ +instance : Inhabited (AdeleRing R K) := ⟨0⟩ -instance : TopologicalSpace (AdeleRing K) := instTopologicalSpaceProd +instance : TopologicalSpace (AdeleRing R K) := instTopologicalSpaceProd -instance : TopologicalRing (AdeleRing K) := instTopologicalRingProd +instance : TopologicalRing (AdeleRing R K) := instTopologicalRingProd -instance : Algebra K (AdeleRing K) := Prod.algebra _ _ _ +instance : Algebra K (AdeleRing R K) := Prod.algebra _ _ _ @[simp] theorem algebraMap_fst_apply (x : K) (v : InfinitePlace K) : - (algebraMap K (AdeleRing K) x).1 v = x := rfl + (algebraMap K (AdeleRing R K) x).1 v = x := rfl @[simp] -theorem algebraMap_snd_apply (x : K) (v : HeightOneSpectrum (𝓞 K)) : - (algebraMap K (AdeleRing K) x).2 v = x := rfl +theorem algebraMap_snd_apply (x : K) (v : HeightOneSpectrum R) : + (algebraMap K (AdeleRing R K) x).2 v = x := rfl -theorem algebraMap_injective : Function.Injective (algebraMap K (AdeleRing K)) := +theorem algebraMap_injective [NumberField K] : Function.Injective (algebraMap K (AdeleRing R K)) := fun _ _ hxy => (algebraMap K _).injective (Prod.ext_iff.1 hxy).1 /-- The subgroup of principal adeles `(x)ᵥ` where `x ∈ K`. -/ -abbrev principalSubgroup : AddSubgroup (AdeleRing K) := (algebraMap K _).range.toAddSubgroup +abbrev principalSubgroup : AddSubgroup (AdeleRing R K) := (algebraMap K _).range.toAddSubgroup end AdeleRing diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean index 3b10a1368dea7..46dd10425cc97 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean @@ -388,7 +388,7 @@ theorem nnnorm_eq_sup_normAtPlace (x : mixedSpace K) : (univ.image (fun w : {w : InfinitePlace K // IsComplex w} ↦ w.1)) := by ext; simp [isReal_or_isComplex] rw [this, sup_union, univ.sup_image, univ.sup_image, - Prod.nnnorm_def', Pi.nnnorm_def, Pi.nnnorm_def] + Prod.nnnorm_def, Pi.nnnorm_def, Pi.nnnorm_def] congr · ext w simp [normAtPlace_apply_isReal w.prop] @@ -979,6 +979,7 @@ theorem neg_of_mem_negA_plusPart (hx : x ∈ negAt s '' (plusPart A)) {w : {w // rw [negAt_apply_of_isReal_and_not_mem _ hw] exact hy.2 w +open scoped Function in -- required for scoped `on` notation /-- The images of `plusPart` by `negAt` are pairwise disjoint. -/ theorem disjoint_negAt_plusPart : Pairwise (Disjoint on (fun s ↦ negAt s '' (plusPart A))) := by intro s t hst diff --git a/Mathlib/NumberTheory/NumberField/Embeddings.lean b/Mathlib/NumberTheory/NumberField/Embeddings.lean index 1d5de4b41ece5..c057b1aa09be8 100644 --- a/Mathlib/NumberTheory/NumberField/Embeddings.lean +++ b/Mathlib/NumberTheory/NumberField/Embeddings.lean @@ -1096,4 +1096,28 @@ lemma infinitePlace_apply (v : InfinitePlace ℚ) (x : ℚ) : v x = |x| := by instance : Subsingleton (InfinitePlace ℚ) where allEq a b := by ext; simp +lemma isReal_infinitePlace : InfinitePlace.IsReal (infinitePlace) := + ⟨Rat.castHom ℂ, by ext; simp, rfl⟩ + end Rat + +/- + +## Totally real number fields + +-/ + +namespace NumberField + +/-- A number field `K` is totally real if all of its infinite places +are real. In other words, the image of every ring homomorphism `K → ℂ` +is a subset of `ℝ`. -/ +class IsTotallyReal (K : Type*) [Field K] [NumberField K] where + isReal : ∀ v : InfinitePlace K, v.IsReal + +instance : IsTotallyReal ℚ where + isReal v := by + rw [Subsingleton.elim v Rat.infinitePlace] + exact Rat.isReal_infinitePlace + +end NumberField diff --git a/Mathlib/NumberTheory/Padics/PadicNumbers.lean b/Mathlib/NumberTheory/Padics/PadicNumbers.lean index a2b72fdf7bbad..c03112d1c05ed 100644 --- a/Mathlib/NumberTheory/Padics/PadicNumbers.lean +++ b/Mathlib/NumberTheory/Padics/PadicNumbers.lean @@ -977,10 +977,9 @@ lemma valuation_intCast (n : ℤ) : valuation (n : ℚ_[p]) = padicValInt p n := lemma valuation_natCast (n : ℕ) : valuation (n : ℚ_[p]) = padicValNat p n := by rw [← Rat.cast_natCast, valuation_ratCast, padicValRat.of_nat] --- See note [no_index around OfNat.ofNat] @[simp] lemma valuation_ofNat (n : ℕ) [n.AtLeastTwo] : - valuation (no_index (OfNat.ofNat n : ℚ_[p])) = padicValNat p n := + valuation (ofNat(n) : ℚ_[p]) = padicValNat p n := valuation_natCast n @[simp] diff --git a/Mathlib/NumberTheory/Transcendental/Liouville/Residual.lean b/Mathlib/NumberTheory/Transcendental/Liouville/Residual.lean index 9e0eff961d50c..fc253159650ba 100644 --- a/Mathlib/NumberTheory/Transcendental/Liouville/Residual.lean +++ b/Mathlib/NumberTheory/Transcendental/Liouville/Residual.lean @@ -34,7 +34,6 @@ theorem IsGδ.setOf_liouville : IsGδ { x | Liouville x } := by refine isOpen_iUnion fun a => isOpen_iUnion fun b => isOpen_iUnion fun _hb => ?_ exact isOpen_ball.inter isClosed_singleton.isOpen_compl -@[deprecated (since := "2024-02-15")] alias isGδ_setOf_liouville := IsGδ.setOf_liouville theorem setOf_liouville_eq_irrational_inter_iInter_iUnion : { x | Liouville x } = diff --git a/Mathlib/NumberTheory/Zsqrtd/Basic.lean b/Mathlib/NumberTheory/Zsqrtd/Basic.lean index 9d477daf4ffca..5165598a7e5fe 100644 --- a/Mathlib/NumberTheory/Zsqrtd/Basic.lean +++ b/Mathlib/NumberTheory/Zsqrtd/Basic.lean @@ -230,7 +230,7 @@ theorem natCast_re (n : ℕ) : (n : ℤ√d).re = n := rfl @[simp] -theorem ofNat_re (n : ℕ) [n.AtLeastTwo] : (no_index (OfNat.ofNat n) : ℤ√d).re = n := +theorem ofNat_re (n : ℕ) [n.AtLeastTwo] : (ofNat(n) : ℤ√d).re = n := rfl @[simp] @@ -238,7 +238,7 @@ theorem natCast_im (n : ℕ) : (n : ℤ√d).im = 0 := rfl @[simp] -theorem ofNat_im (n : ℕ) [n.AtLeastTwo] : (no_index (OfNat.ofNat n) : ℤ√d).im = 0 := +theorem ofNat_im (n : ℕ) [n.AtLeastTwo] : (ofNat(n) : ℤ√d).im = 0 := rfl theorem natCast_val (n : ℕ) : (n : ℤ√d) = ⟨n, 0⟩ := @@ -927,9 +927,7 @@ def lift {d : ℤ} : { r : R // r * r = ↑d } ≃ (ℤ√d →+* R) where simp only [mul_re, Int.cast_add, Int.cast_mul, mul_im, this, r.prop] ring } invFun f := ⟨f sqrtd, by rw [← f.map_mul, dmuld, map_intCast]⟩ - left_inv r := by - ext - simp + left_inv r := by simp right_inv f := by ext simp diff --git a/Mathlib/Order/BooleanSubalgebra.lean b/Mathlib/Order/BooleanSubalgebra.lean index 24751706b1099..bf798efa14b9a 100644 --- a/Mathlib/Order/BooleanSubalgebra.lean +++ b/Mathlib/Order/BooleanSubalgebra.lean @@ -160,10 +160,10 @@ instance instTop : Top (BooleanSubalgebra α) where /-- The trivial boolean subalgebra of a lattice. -/ instance instBot : Bot (BooleanSubalgebra α) where bot.carrier := {⊥, ⊤} - bot.bot_mem' := by aesop - bot.compl_mem' := by aesop - bot.supClosed' _ := by aesop - bot.infClosed' _ := by aesop + bot.bot_mem' := by simp + bot.compl_mem' := by simp + bot.supClosed' _ := by simp + bot.infClosed' _ := by simp /-- The inf of two boolean subalgebras is their intersection. -/ instance instInf : Min (BooleanSubalgebra α) where @@ -344,6 +344,12 @@ lemma subset_closure : s ⊆ closure s := fun _ hx ↦ mem_closure.2 fun _ hK lemma closure_mono (hst : s ⊆ t) : closure s ≤ closure t := sInf_le_sInf fun _L ↦ hst.trans +lemma latticeClosure_subset_closure : latticeClosure s ⊆ closure s := + latticeClosure_min subset_closure (closure s).isSublattice + +@[simp] lemma closure_latticeClosure (s : Set α) : closure (latticeClosure s) = closure s := + le_antisymm (closure_le.2 latticeClosure_subset_closure) (closure_mono subset_latticeClosure) + /-- An induction principle for closure membership. If `p` holds for `⊥` and all elements of `s`, and is preserved under suprema and complement, then `p` holds for all elements of the closure of `s`. -/ @[elab_as_elim] @@ -363,8 +369,8 @@ lemma closure_bot_sup_induction {p : ∀ g ∈ closure s, Prop} (mem : ∀ x hx, section sdiff_sup -variable (sup : SupClosed s) (inf : InfClosed s) (bot_mem : ⊥ ∈ s) (top_mem : ⊤ ∈ s) -include sup inf bot_mem top_mem +variable (isSublattice : IsSublattice s) (bot_mem : ⊥ ∈ s) (top_mem : ⊤ ∈ s) +include isSublattice bot_mem top_mem theorem mem_closure_iff_sup_sdiff {a : α} : a ∈ closure s ↔ ∃ t : Finset (s × s), a = t.sup fun x ↦ x.1.1 \ x.2.1 := by @@ -381,21 +387,21 @@ theorem mem_closure_iff_sup_sdiff {a : α} : simp_rw [Finset.sup_insert, compl_sup, eq] refine tc.induction ⟨∅, by simp⟩ fun ⟨z, w⟩ tc _ ⟨t, eq⟩ ↦ ?_ simp_rw [Finset.sup_insert, inf_sup_left, eq] - refine ⟨{(z, ⟨_, sup x.2 w.2⟩), (⟨_, inf y.2 z.2⟩, w)} ∪ t, ?_⟩ + use {(z, ⟨_, isSublattice.supClosed x.2 w.2⟩), (⟨_, isSublattice.infClosed y.2 z.2⟩, w)} ∪ t simp_rw [Finset.sup_union, Finset.sup_insert, Finset.sup_singleton, sdiff_eq, compl_sup, inf_left_comm z.1, compl_inf, compl_compl, inf_sup_right, inf_assoc] @[elab_as_elim] theorem closure_sdiff_sup_induction {p : ∀ g ∈ closure s, Prop} (sdiff : ∀ x hx y hy, p (x \ y) (sdiff_mem (subset_closure hx) (subset_closure hy))) - (sup' : ∀ x hx y hy, p x hx → p y hy → p (x ⊔ y) (sup_mem hx hy)) - {x} (hx : x ∈ closure s) : p x hx := by - obtain ⟨t, rfl⟩ := (mem_closure_iff_sup_sdiff sup inf bot_mem top_mem).mp hx + (sup : ∀ x hx y hy, p x hx → p y hy → p (x ⊔ y) (sup_mem hx hy)) + (x) (hx : x ∈ closure s) : p x hx := by + obtain ⟨t, rfl⟩ := (mem_closure_iff_sup_sdiff isSublattice bot_mem top_mem).mp hx revert hx classical refine t.induction (by simpa using sdiff _ bot_mem _ bot_mem) fun x t _ ih hxt ↦ ?_ simp only [Finset.sup_insert] at hxt ⊢ - exact sup' _ _ _ _ (sdiff _ x.1.2 _ x.2.2) - (ih <| (mem_closure_iff_sup_sdiff sup inf bot_mem top_mem).mpr ⟨_, rfl⟩) + exact sup _ _ _ ((mem_closure_iff_sup_sdiff isSublattice bot_mem top_mem).mpr ⟨_, rfl⟩) + (sdiff _ x.1.2 _ x.2.2) (ih _) end sdiff_sup diff --git a/Mathlib/Order/CompleteLattice.lean b/Mathlib/Order/CompleteLattice.lean index c225c90223d8d..d1c676bce4fdf 100644 --- a/Mathlib/Order/CompleteLattice.lean +++ b/Mathlib/Order/CompleteLattice.lean @@ -1726,11 +1726,11 @@ theorem iInf_sup_iInf_le (f g : ι → α) : (⨅ i, f i) ⊔ ⨅ i, g i ≤ ⨅ theorem disjoint_sSup_left {a : Set α} {b : α} (d : Disjoint (sSup a) b) {i} (hi : i ∈ a) : Disjoint i b := - disjoint_iff_inf_le.mpr (iSup₂_le_iff.1 (iSup_inf_le_sSup_inf.trans d.le_bot) i hi : _) + disjoint_iff_inf_le.mpr (iSup₂_le_iff.1 (iSup_inf_le_sSup_inf.trans d.le_bot) i hi :) theorem disjoint_sSup_right {a : Set α} {b : α} (d : Disjoint b (sSup a)) {i} (hi : i ∈ a) : Disjoint b i := - disjoint_iff_inf_le.mpr (iSup₂_le_iff.mp (iSup_inf_le_inf_sSup.trans d.le_bot) i hi : _) + disjoint_iff_inf_le.mpr (iSup₂_le_iff.mp (iSup_inf_le_inf_sSup.trans d.le_bot) i hi :) end CompleteLattice diff --git a/Mathlib/Order/Disjointed.lean b/Mathlib/Order/Disjointed.lean index 766108a6dec8c..1366a858c3dcd 100644 --- a/Mathlib/Order/Disjointed.lean +++ b/Mathlib/Order/Disjointed.lean @@ -37,6 +37,8 @@ Related to the TODO in the module docstring of `Mathlib.Order.PartialSups`. variable {α : Type*} +open scoped Function -- required for scoped `on` notation + section GeneralizedBooleanAlgebra variable [GeneralizedBooleanAlgebra α] diff --git a/Mathlib/Order/Filter/Bases.lean b/Mathlib/Order/Filter/Bases.lean index 42dfb3d129084..9aeebbf245242 100644 --- a/Mathlib/Order/Filter/Bases.lean +++ b/Mathlib/Order/Filter/Bases.lean @@ -578,6 +578,7 @@ theorem _root_.Disjoint.exists_mem_filter_basis (h : Disjoint l l') (hl : l.HasB (hl' : l'.HasBasis p' s') : ∃ i, p i ∧ ∃ i', p' i' ∧ Disjoint (s i) (s' i') := (hl.disjoint_iff hl').1 h +open scoped Function in -- required for scoped `on` notation theorem _root_.Pairwise.exists_mem_filter_basis_of_disjoint {I} [Finite I] {l : I → Filter α} {ι : I → Sort*} {p : ∀ i, ι i → Prop} {s : ∀ i, ι i → Set α} (hd : Pairwise (Disjoint on l)) (h : ∀ i, (l i).HasBasis (p i) (s i)) : diff --git a/Mathlib/Order/Filter/Germ/Basic.lean b/Mathlib/Order/Filter/Germ/Basic.lean index fe099e6b204fa..f02186b5a0c53 100644 --- a/Mathlib/Order/Filter/Germ/Basic.lean +++ b/Mathlib/Order/Filter/Germ/Basic.lean @@ -427,16 +427,14 @@ theorem natCast_def [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] : - ((no_index (OfNat.ofNat n : α → M)) : Germ l M) = OfNat.ofNat n := + ((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] : - ((no_index (OfNat.ofNat n : M)) : Germ l M) = OfNat.ofNat n := + ((ofNat(n) : M) : Germ l M) = OfNat.ofNat n := rfl instance instIntCast [IntCast M] : IntCast (Germ l M) where intCast n := (n : α → M) diff --git a/Mathlib/Order/Filter/Ultrafilter.lean b/Mathlib/Order/Filter/Ultrafilter.lean index d171f958a9ca2..f64833156e58c 100644 --- a/Mathlib/Order/Filter/Ultrafilter.lean +++ b/Mathlib/Order/Filter/Ultrafilter.lean @@ -261,7 +261,7 @@ theorem comap_pure {m : α → β} (a : α) (inj : Injective m) (large) : rw [coe_pure, ← principal_singleton, ← image_singleton, preimage_image_eq _ inj] theorem pure_injective : Injective (pure : α → Ultrafilter α) := fun _ _ h => - Filter.pure_injective (congr_arg Ultrafilter.toFilter h : _) + Filter.pure_injective (congr_arg Ultrafilter.toFilter h :) instance [Inhabited α] : Inhabited (Ultrafilter α) := ⟨pure default⟩ diff --git a/Mathlib/Order/Height.lean b/Mathlib/Order/Height.lean index 22b9697c5dd05..be2eaa7acd45b 100644 --- a/Mathlib/Order/Height.lean +++ b/Mathlib/Order/Height.lean @@ -291,10 +291,10 @@ theorem chainHeight_union_le : (s ∪ t).chainHeight ≤ s.chainHeight + t.chain let l₂ := l.filter (· ∈ t) have hl₁ : ↑l₁.length ≤ s.chainHeight := by apply Set.length_le_chainHeight_of_mem_subchain - exact ⟨hl.1.sublist (filter_sublist _), fun i h ↦ by simpa using (of_mem_filter h : _)⟩ + exact ⟨hl.1.sublist (filter_sublist _), fun i h ↦ by simpa using (of_mem_filter h :)⟩ have hl₂ : ↑l₂.length ≤ t.chainHeight := by apply Set.length_le_chainHeight_of_mem_subchain - exact ⟨hl.1.sublist (filter_sublist _), fun i h ↦ by simpa using (of_mem_filter h : _)⟩ + exact ⟨hl.1.sublist (filter_sublist _), fun i h ↦ by simpa using (of_mem_filter h :)⟩ refine le_trans ?_ (add_le_add hl₁ hl₂) simp_rw [l₁, l₂, ← Nat.cast_add, ← Multiset.coe_card, ← Multiset.card_add, ← Multiset.filter_coe] diff --git a/Mathlib/Order/Hom/Order.lean b/Mathlib/Order/Hom/Order.lean index 1fdadf691bf97..730ef64eb8a36 100644 --- a/Mathlib/Order/Hom/Order.lean +++ b/Mathlib/Order/Hom/Order.lean @@ -78,11 +78,11 @@ instance orderTop [Preorder β] [OrderTop β] : OrderTop (α →o β) where le_top _ _ := le_top instance [CompleteLattice β] : InfSet (α →o β) where - sInf s := ⟨fun x => ⨅ f ∈ s, (f : _) x, fun _ _ h => iInf₂_mono fun f _ => f.mono h⟩ + sInf s := ⟨fun x => ⨅ f ∈ s, (f :) x, fun _ _ h => iInf₂_mono fun f _ => f.mono h⟩ @[simp] theorem sInf_apply [CompleteLattice β] (s : Set (α →o β)) (x : α) : - sInf s x = ⨅ f ∈ s, (f : _) x := + sInf s x = ⨅ f ∈ s, (f :) x := rfl theorem iInf_apply {ι : Sort*} [CompleteLattice β] (f : ι → α →o β) (x : α) : @@ -95,11 +95,11 @@ theorem coe_iInf {ι : Sort*} [CompleteLattice β] (f : ι → α →o β) : funext x; simp [iInf_apply] instance [CompleteLattice β] : SupSet (α →o β) where - sSup s := ⟨fun x => ⨆ f ∈ s, (f : _) x, fun _ _ h => iSup₂_mono fun f _ => f.mono h⟩ + sSup s := ⟨fun x => ⨆ f ∈ s, (f :) x, fun _ _ h => iSup₂_mono fun f _ => f.mono h⟩ @[simp] theorem sSup_apply [CompleteLattice β] (s : Set (α →o β)) (x : α) : - sSup s x = ⨆ f ∈ s, (f : _) x := + sSup s x = ⨆ f ∈ s, (f :) x := rfl theorem iSup_apply {ι : Sort*} [CompleteLattice β] (f : ι → α →o β) (x : α) : diff --git a/Mathlib/Order/InitialSeg.lean b/Mathlib/Order/InitialSeg.lean index 9d83ed92d1c08..d00f90a854fe1 100644 --- a/Mathlib/Order/InitialSeg.lean +++ b/Mathlib/Order/InitialSeg.lean @@ -6,6 +6,7 @@ Authors: Mario Carneiro, Floris van Doorn, Violeta Hernández Palacios import Mathlib.Data.Sum.Order import Mathlib.Logic.Equiv.Set import Mathlib.Order.RelIso.Set +import Mathlib.Order.UpperLower.Basic import Mathlib.Order.WellFounded /-! @@ -670,6 +671,9 @@ alias ⟨_, map_isMin⟩ := isMin_apply_iff theorem map_bot [PartialOrder α] [OrderBot α] [OrderBot β] (f : α ≤i β) : f ⊥ = ⊥ := (map_isMin f isMin_bot).eq_bot +theorem image_Iio [PartialOrder α] (f : α ≤i β) (a : α) : f '' Set.Iio a = Set.Iio (f a) := + f.toOrderEmbedding.image_Iio f.isLowerSet_range a + theorem le_apply_iff [LinearOrder α] (f : α ≤i β) : b ≤ f a ↔ ∃ c ≤ a, f c = b := by constructor · intro h @@ -726,6 +730,9 @@ alias ⟨_, map_isMin⟩ := isMin_apply_iff theorem map_bot [PartialOrder α] [OrderBot α] [OrderBot β] (f : α (ne_of_mem_of_not_mem hi hx : _)) (ht x) + Disjoint.mono_right (biSup_mono fun _ hi => (ne_of_mem_of_not_mem hi hx :)) (ht x) @[deprecated (since := "2024-11-24")] alias CompleteLattice.Independent.disjoint_biSup := iSupIndep.disjoint_biSup @@ -565,6 +565,7 @@ alias setIndependent_iff_pairwiseDisjoint := sSupIndep_iff_pairwiseDisjoint alias ⟨_, _root_.Set.PairwiseDisjoint.sSupIndep⟩ := sSupIndep_iff_pairwiseDisjoint +open scoped Function in -- required for scoped `on` notation theorem iSupIndep_iff_pairwiseDisjoint {f : ι → α} : iSupIndep f ↔ Pairwise (Disjoint on f) := ⟨iSupIndep.pairwiseDisjoint, fun hs _ => disjoint_iSup_iff.2 fun _ => disjoint_iSup_iff.2 fun hij => hs hij.symm⟩ diff --git a/Mathlib/Order/WellFounded.lean b/Mathlib/Order/WellFounded.lean index 17b5c843b5ec7..5aa9a2eaa166b 100644 --- a/Mathlib/Order/WellFounded.lean +++ b/Mathlib/Order/WellFounded.lean @@ -37,6 +37,7 @@ instance : IsIrrefl α WellFoundedRelation.rel := IsAsymm.isIrrefl theorem mono (hr : WellFounded r) (h : ∀ a b, r' a b → r a b) : WellFounded r' := Subrelation.wf (h _ _) hr +open scoped Function in -- required for scoped `on` notation theorem onFun {α β : Sort*} {r : β → β → Prop} {f : α → β} : WellFounded r → WellFounded (r on f) := InvImage.wf _ diff --git a/Mathlib/Order/WellFoundedSet.lean b/Mathlib/Order/WellFoundedSet.lean index c2b623cec5ed6..b021f89341c75 100644 --- a/Mathlib/Order/WellFoundedSet.lean +++ b/Mathlib/Order/WellFoundedSet.lean @@ -46,6 +46,8 @@ Prove that `s` is partial well ordered iff it has no infinite descending chain o assert_not_exists OrderedSemiring +open scoped Function -- required for scoped `on` notation + variable {ι α β γ : Type*} {π : ι → Type*} namespace Set diff --git a/Mathlib/Probability/Independence/Basic.lean b/Mathlib/Probability/Independence/Basic.lean index a7ef0fcd188d8..3bdcf8aaffe0c 100644 --- a/Mathlib/Probability/Independence/Basic.lean +++ b/Mathlib/Probability/Independence/Basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ import Mathlib.Probability.Independence.Kernel +import Mathlib.MeasureTheory.Measure.Prod /-! # Independence of sets of sets and measure spaces (σ-algebras) diff --git a/Mathlib/Probability/Independence/Kernel.lean b/Mathlib/Probability/Independence/Kernel.lean index c83e793c26c67..1b8fc5fb8bb20 100644 --- a/Mathlib/Probability/Independence/Kernel.lean +++ b/Mathlib/Probability/Independence/Kernel.lean @@ -3,10 +3,10 @@ Copyright (c) 2023 Rémy Degenne. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ -import Mathlib.MeasureTheory.Constructions.Pi import Mathlib.Probability.ConditionalProbability import Mathlib.Probability.Kernel.Basic import Mathlib.Tactic.Peel +import Mathlib.MeasureTheory.MeasurableSpace.Pi /-! # Independence with respect to a kernel and a measure diff --git a/Mathlib/Probability/Kernel/Composition/Basic.lean b/Mathlib/Probability/Kernel/Composition/Basic.lean index 30d4f488d8861..c5cdd7bb8284a 100644 --- a/Mathlib/Probability/Kernel/Composition/Basic.lean +++ b/Mathlib/Probability/Kernel/Composition/Basic.lean @@ -77,6 +77,7 @@ We define a kernel composition-product `compProd : Kernel α β → Kernel (α × β) γ → Kernel α (β × γ)`. -/ +open scoped Function -- required for scoped `on` notation variable {γ : Type*} {mγ : MeasurableSpace γ} {s : Set (β × γ)} @@ -308,7 +309,7 @@ theorem ae_null_of_compProd_null (h : (κ ⊗ₖ η) a s = 0) : rw [Filter.eventuallyLE_antisymm_iff] exact ⟨Filter.EventuallyLE.trans_eq - (Filter.Eventually.of_forall fun x => (measure_mono (Set.preimage_mono hst) : _)) ht, + (Filter.Eventually.of_forall fun x => measure_mono (Set.preimage_mono hst)) ht, Filter.Eventually.of_forall fun x => zero_le _⟩ theorem ae_ae_of_ae_compProd {p : β × γ → Prop} (h : ∀ᵐ bc ∂(κ ⊗ₖ η) a, p bc) : @@ -1063,6 +1064,104 @@ lemma snd_swapRight (κ : Kernel α (β × γ)) : snd (swapRight κ) = fst κ := end FstSnd +section sectLsectR + +variable {γ δ : Type*} {mγ : MeasurableSpace γ} {mδ : MeasurableSpace δ} + +/-- Define a `Kernel α γ` from a `Kernel (α × β) γ` by taking the comap of `fun a ↦ (a, b)` for +a given `b : β`. -/ +noncomputable def sectL (κ : Kernel (α × β) γ) (b : β) : Kernel α γ := + comap κ (fun a ↦ (a, b)) (measurable_id.prod_mk measurable_const) + +@[simp] theorem sectL_apply (κ : Kernel (α × β) γ) (b : β) (a : α) : sectL κ b a = κ (a, b) := rfl + +@[simp] lemma sectL_zero (b : β) : sectL (0 : Kernel (α × β) γ) b = 0 := by simp [sectL] + +instance (κ : Kernel (α × β) γ) (b : β) [IsMarkovKernel κ] : IsMarkovKernel (sectL κ b) := by + rw [sectL]; infer_instance + +instance (κ : Kernel (α × β) γ) (b : β) [IsZeroOrMarkovKernel κ] : + IsZeroOrMarkovKernel (sectL κ b) := by + rw [sectL]; infer_instance + +instance (κ : Kernel (α × β) γ) (b : β) [IsFiniteKernel κ] : IsFiniteKernel (sectL κ b) := by + rw [sectL]; infer_instance + +instance (κ : Kernel (α × β) γ) (b : β) [IsSFiniteKernel κ] : IsSFiniteKernel (sectL κ b) := by + rw [sectL]; infer_instance + +instance (κ : Kernel (α × β) γ) (a : α) (b : β) [NeZero (κ (a, b))] : NeZero ((sectL κ b) a) := by + rw [sectL_apply]; infer_instance + +instance (priority := 100) {κ : Kernel (α × β) γ} [∀ b, IsMarkovKernel (sectL κ b)] : + IsMarkovKernel κ := by + refine ⟨fun _ ↦ ⟨?_⟩⟩ + rw [← sectL_apply, measure_univ] + +--I'm not sure this lemma is actually useful +lemma comap_sectL (κ : Kernel (α × β) γ) (b : β) {f : δ → α} (hf : Measurable f) : + comap (sectL κ b) f hf = comap κ (fun d ↦ (f d, b)) (hf.prod_mk measurable_const) := by + ext d s + rw [comap_apply, sectL_apply, comap_apply] + +@[simp] +lemma sectL_prodMkLeft (α : Type*) [MeasurableSpace α] (κ : Kernel β γ) (a : α) {b : β} : + sectL (prodMkLeft α κ) b a = κ b := rfl + +@[simp] +lemma sectL_prodMkRight (β : Type*) [MeasurableSpace β] (κ : Kernel α γ) (b : β) : + sectL (prodMkRight β κ) b = κ := rfl + +/-- Define a `Kernel β γ` from a `Kernel (α × β) γ` by taking the comap of `fun b ↦ (a, b)` for +a given `a : α`. -/ +noncomputable def sectR (κ : Kernel (α × β) γ) (a : α) : Kernel β γ := + comap κ (fun b ↦ (a, b)) (measurable_const.prod_mk measurable_id) + +@[simp] theorem sectR_apply (κ : Kernel (α × β) γ) (b : β) (a : α) : sectR κ a b = κ (a, b) := rfl + +@[simp] lemma sectR_zero (a : α) : sectR (0 : Kernel (α × β) γ) a = 0 := by simp [sectR] + +instance (κ : Kernel (α × β) γ) (a : α) [IsMarkovKernel κ] : IsMarkovKernel (sectR κ a) := by + rw [sectR]; infer_instance + +instance (κ : Kernel (α × β) γ) (a : α) [IsZeroOrMarkovKernel κ] : + IsZeroOrMarkovKernel (sectR κ a) := by + rw [sectR]; infer_instance + +instance (κ : Kernel (α × β) γ) (a : α) [IsFiniteKernel κ] : IsFiniteKernel (sectR κ a) := by + rw [sectR]; infer_instance + +instance (κ : Kernel (α × β) γ) (a : α) [IsSFiniteKernel κ] : IsSFiniteKernel (sectR κ a) := by + rw [sectR]; infer_instance + +instance (κ : Kernel (α × β) γ) (a : α) (b : β) [NeZero (κ (a, b))] : NeZero ((sectR κ a) b) := by + rw [sectR_apply]; infer_instance + +instance (priority := 100) {κ : Kernel (α × β) γ} [∀ b, IsMarkovKernel (sectR κ b)] : + IsMarkovKernel κ := by + refine ⟨fun _ ↦ ⟨?_⟩⟩ + rw [← sectR_apply, measure_univ] + +--I'm not sure this lemma is actually useful +lemma comap_sectR (κ : Kernel (α × β) γ) (a : α) {f : δ → β} (hf : Measurable f) : + comap (sectR κ a) f hf = comap κ (fun d ↦ (a, f d)) (measurable_const.prod_mk hf) := by + ext d s + rw [comap_apply, sectR_apply, comap_apply] + +@[simp] +lemma sectR_prodMkLeft (α : Type*) [MeasurableSpace α] (κ : Kernel β γ) (a : α) : + sectR (prodMkLeft α κ) a = κ := rfl + +@[simp] +lemma sectR_prodMkRight (β : Type*) [MeasurableSpace β] (κ : Kernel α γ) (b : β) {a : α} : + sectR (prodMkRight β κ) a b = κ a := rfl + +@[simp] lemma sectL_swapRight (κ : Kernel (α × β) γ) : sectL (swapLeft κ) = sectR κ := rfl + +@[simp] lemma sectR_swapRight (κ : Kernel (α × β) γ) : sectR (swapLeft κ) = sectL κ := rfl + +end sectLsectR + section Comp /-! ### Composition of two kernels -/ @@ -1204,6 +1303,7 @@ variable {γ δ : Type*} {mγ : MeasurableSpace γ} {mδ : MeasurableSpace δ} noncomputable def prod (κ : Kernel α β) (η : Kernel α γ) : Kernel α (β × γ) := κ ⊗ₖ swapLeft (prodMkLeft β η) +@[inherit_doc] scoped[ProbabilityTheory] infixl:100 " ×ₖ " => ProbabilityTheory.Kernel.prod theorem prod_apply' (κ : Kernel α β) [IsSFiniteKernel κ] (η : Kernel α γ) [IsSFiniteKernel η] diff --git a/Mathlib/Probability/Kernel/Composition/IntegralCompProd.lean b/Mathlib/Probability/Kernel/Composition/IntegralCompProd.lean index 0f7fef4ed3332..a9ab89e8cf8f4 100644 --- a/Mathlib/Probability/Kernel/Composition/IntegralCompProd.lean +++ b/Mathlib/Probability/Kernel/Composition/IntegralCompProd.lean @@ -134,7 +134,7 @@ theorem _root_.MeasureTheory.Integrable.integral_compProd [NormedSpace ℝ E] (norm_integral_le_integral_norm _).trans_eq <| (norm_of_nonneg <| integral_nonneg_of_ae <| - Eventually.of_forall fun y => (norm_nonneg (f (x, y)) : _)).symm + Eventually.of_forall fun y => (norm_nonneg (f (x, y)) :)).symm /-! ### Bochner integral -/ diff --git a/Mathlib/Probability/Kernel/Composition/MeasureCompProd.lean b/Mathlib/Probability/Kernel/Composition/MeasureCompProd.lean index b57ca2ef11459..64933b8111dc2 100644 --- a/Mathlib/Probability/Kernel/Composition/MeasureCompProd.lean +++ b/Mathlib/Probability/Kernel/Composition/MeasureCompProd.lean @@ -73,6 +73,14 @@ lemma compProd_apply_prod [SFinite μ] [IsSFiniteKernel κ] rw [indicator_apply] split_ifs with ha <;> simp [ha] +lemma _root_.ProbabilityTheory.Kernel.compProd_apply_eq_compProd_sectR {γ : Type*} + {mγ : MeasurableSpace γ} (κ : Kernel α β) (η : Kernel (α × β) γ) + [IsSFiniteKernel κ] [IsSFiniteKernel η] (a : α) : + (κ ⊗ₖ η) a = (κ a) ⊗ₘ (Kernel.sectR η a) := by + ext s hs + simp_rw [Kernel.compProd_apply hs, compProd_apply hs, Kernel.sectR_apply] + rfl + lemma compProd_congr [IsSFiniteKernel κ] [IsSFiniteKernel η] (h : κ =ᵐ[μ] η) : μ ⊗ₘ κ = μ ⊗ₘ η := by by_cases hμ : SFinite μ diff --git a/Mathlib/Probability/Kernel/Disintegration/CDFToKernel.lean b/Mathlib/Probability/Kernel/Disintegration/CDFToKernel.lean index 3dceea49cb6c1..18477b6312708 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CDFToKernel.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CDFToKernel.lean @@ -576,6 +576,7 @@ lemma setLIntegral_toKernel_prod [IsFiniteKernel κ] (hf : IsCondKernelCDF f κ @[deprecated (since := "2024-06-29")] alias set_lintegral_toKernel_prod := setLIntegral_toKernel_prod +open scoped Function in -- required for scoped `on` notation lemma lintegral_toKernel_mem [IsFiniteKernel κ] (hf : IsCondKernelCDF f κ ν) (a : α) {s : Set (β × ℝ)} (hs : MeasurableSet s) : ∫⁻ b, hf.toKernel f (a, b) {y | (b, y) ∈ s} ∂(ν a) = κ a s := by diff --git a/Mathlib/Probability/Kernel/Disintegration/Density.lean b/Mathlib/Probability/Kernel/Disintegration/Density.lean index 30f8ae52d6d41..9dd84c89fa298 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Density.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Density.lean @@ -241,6 +241,7 @@ lemma setIntegral_densityProcess_of_mem (hκν : fst κ ≤ ν) [hν : IsFiniteK @[deprecated (since := "2024-04-17")] alias set_integral_densityProcess_of_mem := setIntegral_densityProcess_of_mem +open scoped Function in -- required for scoped `on` notation lemma setIntegral_densityProcess (hκν : fst κ ≤ ν) [IsFiniteKernel ν] (n : ℕ) (a : α) {s : Set β} (hs : MeasurableSet s) {A : Set γ} (hA : MeasurableSet[countableFiltration γ n] A) : diff --git a/Mathlib/Probability/Moments.lean b/Mathlib/Probability/Moments.lean index b6c17d414699d..7e8db5fd848ce 100644 --- a/Mathlib/Probability/Moments.lean +++ b/Mathlib/Probability/Moments.lean @@ -59,11 +59,18 @@ theorem moment_zero (hp : p ≠ 0) : moment 0 p μ = 0 := by simp only [moment, hp, zero_pow, Ne, not_false_iff, Pi.zero_apply, integral_const, smul_eq_mul, mul_zero, integral_zero] +@[simp] +lemma moment_zero_measure : moment X p (0 : Measure Ω) = 0 := by simp [moment] + @[simp] theorem centralMoment_zero (hp : p ≠ 0) : centralMoment 0 p μ = 0 := by simp only [centralMoment, hp, Pi.zero_apply, integral_const, smul_eq_mul, mul_zero, zero_sub, Pi.pow_apply, Pi.neg_apply, neg_zero, zero_pow, Ne, not_false_iff] +@[simp] +lemma centralMoment_zero_measure : centralMoment X p (0 : Measure Ω) = 0 := by + simp [centralMoment] + theorem centralMoment_one' [IsFiniteMeasure μ] (h_int : Integrable X μ) : centralMoment X 1 μ = (1 - (μ Set.univ).toReal) * μ[X] := by simp only [centralMoment, Pi.sub_apply, pow_one] @@ -118,7 +125,6 @@ theorem cgf_zero_measure : cgf X (0 : Measure Ω) t = 0 := by theorem mgf_const' (c : ℝ) : mgf (fun _ => c) μ t = (μ Set.univ).toReal * exp (t * c) := by simp only [mgf, integral_const, smul_eq_mul] --- @[simp] -- Porting note: `simp only` already proves this theorem mgf_const (c : ℝ) [IsProbabilityMeasure μ] : mgf (fun _ => c) μ t = exp (t * c) := by simp only [mgf_const', measure_univ, ENNReal.one_toReal, one_mul] @@ -139,7 +145,6 @@ theorem cgf_const [IsProbabilityMeasure μ] (c : ℝ) : cgf (fun _ => c) μ t = theorem mgf_zero' : mgf X μ 0 = (μ Set.univ).toReal := by simp only [mgf, zero_mul, exp_zero, integral_const, smul_eq_mul, mul_one] --- @[simp] -- Porting note: `simp only` already proves this theorem mgf_zero [IsProbabilityMeasure μ] : mgf X μ 0 = 1 := by simp only [mgf_zero', measure_univ, ENNReal.one_toReal] @@ -180,6 +185,12 @@ theorem mgf_pos [IsProbabilityMeasure μ] (h_int_X : Integrable (fun ω => exp ( 0 < mgf X μ t := mgf_pos' (IsProbabilityMeasure.ne_zero μ) h_int_X +lemma exp_cgf_of_neZero [hμ : NeZero μ] (hX : Integrable (fun ω ↦ exp (t * X ω)) μ) : + exp (cgf X μ t) = mgf X μ t := by rw [cgf, exp_log (mgf_pos' hμ.out hX)] + +lemma exp_cgf [IsProbabilityMeasure μ] (hX : Integrable (fun ω ↦ exp (t * X ω)) μ) : + exp (cgf X μ t) = mgf X μ t := by rw [cgf, exp_log (mgf_pos hX)] + lemma mgf_id_map (hX : AEMeasurable X μ) : mgf id (μ.map X) = mgf X μ := by ext t rw [mgf, integral_map hX] @@ -202,6 +213,8 @@ theorem mgf_const_add (α : ℝ) : mgf (fun ω => α + X ω) μ t = exp (t * α) theorem mgf_add_const (α : ℝ) : mgf (fun ω => X ω + α) μ t = mgf X μ t * exp (t * α) := by simp only [add_comm, mgf_const_add, mul_comm] +section IndepFun + /-- This is a trivial application of `IndepFun.comp` but it will come up frequently. -/ theorem IndepFun.exp_mul {X Y : Ω → ℝ} (h_indep : IndepFun X Y μ) (s t : ℝ) : IndepFun (fun ω => exp (s * X ω)) (fun ω => exp (t * Y ω)) μ := by @@ -302,6 +315,8 @@ theorem iIndepFun.cgf_sum {X : ι → Ω → ℝ} · rw [h_indep.mgf_sum h_meas] · exact (mgf_pos (h_int j hj)).ne' +end IndepFun + theorem mgf_congr_of_identDistrib (X : Ω → ℝ) {Ω' : Type*} {m' : MeasurableSpace Ω'} {μ' : Measure Ω'} (X' : Ω' → ℝ) (hident : IdentDistrib X X' μ μ') (t : ℝ) : @@ -318,6 +333,8 @@ theorem mgf_sum_of_identDistrib exact Finset.prod_eq_pow_card fun i hi => mgf_congr_of_identDistrib (X i) (X j) (hident i hi j hj) t +section Chernoff + /-- **Chernoff bound** on the upper tail of a real random variable. -/ theorem measure_ge_le_exp_mul_mgf [IsFiniteMeasure μ] (ε : ℝ) (ht : 0 ≤ t) (h_int : Integrable (fun ω => exp (t * X ω)) μ) : @@ -367,6 +384,8 @@ theorem measure_le_le_exp_cgf [IsFiniteMeasure μ] (ε : ℝ) (ht : t ≤ 0) rw [exp_add] exact mul_le_mul le_rfl (le_exp_log _) mgf_nonneg (exp_pos _).le +end Chernoff + lemma mgf_dirac {x : ℝ} (hX : μ.map X = .dirac x) (t : ℝ) : mgf X μ t = exp (x * t) := by have : IsProbabilityMeasure (μ.map X) := by rw [hX]; infer_instance rw [← mgf_id_map (.of_map_ne_zero <| IsProbabilityMeasure.ne_zero _), mgf, hX, integral_dirac, diff --git a/Mathlib/Probability/StrongLaw.lean b/Mathlib/Probability/StrongLaw.lean index 734ac6c8622b2..0fadd8a41b7a2 100644 --- a/Mathlib/Probability/StrongLaw.lean +++ b/Mathlib/Probability/StrongLaw.lean @@ -62,6 +62,8 @@ open Set (indicator) open scoped Topology MeasureTheory ProbabilityTheory ENNReal NNReal +open scoped Function -- required for scoped `on` notation + namespace ProbabilityTheory /-! ### Prerequisites on truncations -/ @@ -426,7 +428,7 @@ theorem strong_law_aux1 {c : ℝ} (c_one : 1 < c) {ε : ℝ} (εpos : 0 < ε) : · simp only [mem_sigma, mem_range, filter_congr_decidable, mem_filter, and_imp, Sigma.forall] exact fun a b haN hb ↦ ⟨hb.trans_le <| u_mono <| Nat.le_pred_of_lt haN, haN, hb⟩ - all_goals aesop + all_goals simp _ ≤ ∑ j ∈ range (u (N - 1)), c ^ 5 * (c - 1)⁻¹ ^ 3 / ↑j ^ 2 * Var[Y j] := by apply sum_le_sum fun j hj => ?_ rcases @eq_zero_or_pos _ _ j with (rfl | hj) diff --git a/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean b/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean index ff50887857737..1c8844cd09b48 100644 --- a/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean +++ b/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean @@ -797,7 +797,7 @@ def shortComplexH1Iso : (inhomogeneousCochains A).sc' 0 1 2 ≅ shortComplexH1 A /-- The 1-cocycles of the complex of inhomogeneous cochains of `A` are isomorphic to `oneCocycles A`, which is a simpler type. -/ def isoOneCocycles : cocycles A 1 ≅ ModuleCat.of k (oneCocycles A) := - (inhomogeneousCochains A).cyclesIsoSc' _ _ _ (by aesop) (by aesop) ≪≫ + (inhomogeneousCochains A).cyclesIsoSc' _ _ _ (by simp) (by simp) ≪≫ cyclesMapIso (shortComplexH1Iso A) ≪≫ (shortComplexH1 A).moduleCatCyclesIso lemma isoOneCocycles_hom_comp_subtype : @@ -817,7 +817,7 @@ lemma toCocycles_comp_isoOneCocycles_hom : /-- The 1st group cohomology of `A`, defined as the 1st cohomology of the complex of inhomogeneous cochains, is isomorphic to `oneCocycles A ⧸ oneCoboundaries A`, which is a simpler type. -/ def isoH1 : groupCohomology A 1 ≅ ModuleCat.of k (H1 A) := - (inhomogeneousCochains A).homologyIsoSc' _ _ _ (by aesop) (by aesop) ≪≫ + (inhomogeneousCochains A).homologyIsoSc' _ _ _ (by simp) (by simp) ≪≫ homologyMapIso (shortComplexH1Iso A) ≪≫ (shortComplexH1 A).moduleCatHomologyIso lemma groupCohomologyπ_comp_isoH1_hom : @@ -846,7 +846,7 @@ def shortComplexH2Iso : /-- The 2-cocycles of the complex of inhomogeneous cochains of `A` are isomorphic to `twoCocycles A`, which is a simpler type. -/ def isoTwoCocycles : cocycles A 2 ≅ ModuleCat.of k (twoCocycles A) := - (inhomogeneousCochains A).cyclesIsoSc' _ _ _ (by aesop) (by aesop) ≪≫ + (inhomogeneousCochains A).cyclesIsoSc' _ _ _ (by simp) (by simp) ≪≫ cyclesMapIso (shortComplexH2Iso A) ≪≫ (shortComplexH2 A).moduleCatCyclesIso lemma isoTwoCocycles_hom_comp_subtype : @@ -866,7 +866,7 @@ lemma toCocycles_comp_isoTwoCocycles_hom : /-- The 2nd group cohomology of `A`, defined as the 2nd cohomology of the complex of inhomogeneous cochains, is isomorphic to `twoCocycles A ⧸ twoCoboundaries A`, which is a simpler type. -/ def isoH2 : groupCohomology A 2 ≅ ModuleCat.of k (H2 A) := - (inhomogeneousCochains A).homologyIsoSc' _ _ _ (by aesop) (by aesop) ≪≫ + (inhomogeneousCochains A).homologyIsoSc' _ _ _ (by simp) (by simp) ≪≫ homologyMapIso (shortComplexH2Iso A) ≪≫ (shortComplexH2 A).moduleCatHomologyIso lemma groupCohomologyπ_comp_isoH2_hom : diff --git a/Mathlib/RingTheory/AdicCompletion/Functoriality.lean b/Mathlib/RingTheory/AdicCompletion/Functoriality.lean index 4ac8d7b974ccc..80039d291af17 100644 --- a/Mathlib/RingTheory/AdicCompletion/Functoriality.lean +++ b/Mathlib/RingTheory/AdicCompletion/Functoriality.lean @@ -128,7 +128,7 @@ private theorem adicCompletionAux_val_apply (f : M →ₗ[R] N) {n : ℕ} (x : A def map (f : M →ₗ[R] N) : AdicCompletion I M →ₗ[AdicCompletion I R] AdicCompletion I N where toFun := adicCompletionAux I f - map_add' := by aesop + map_add' := by simp map_smul' r x := by ext n simp only [adicCompletionAux_val_apply, smul_eval, smul_eq_mul, RingHom.id_apply] diff --git a/Mathlib/RingTheory/Artinian/Module.lean b/Mathlib/RingTheory/Artinian/Module.lean index 53f91c83eee98..fb67850d6a654 100644 --- a/Mathlib/RingTheory/Artinian/Module.lean +++ b/Mathlib/RingTheory/Artinian/Module.lean @@ -5,6 +5,7 @@ Authors: Chris Hughes -/ import Mathlib.Data.SetLike.Fintype import Mathlib.Order.Filter.EventuallyConst +import Mathlib.RingTheory.Ideal.Prod import Mathlib.RingTheory.Ideal.Quotient.Operations import Mathlib.RingTheory.Nilpotent.Lemmas import Mathlib.RingTheory.Noetherian.Defs @@ -59,10 +60,10 @@ theorem isArtinian_iff (R M) [Semiring R] [AddCommMonoid M] [Module R M] : IsArt WellFounded (· < · : Submodule R M → Submodule R M → Prop) := isWellFounded_iff _ _ -section +section Semiring variable {R M P N : Type*} -variable [Ring R] [AddCommGroup M] [AddCommGroup P] [AddCommGroup N] +variable [Semiring R] [AddCommMonoid M] [AddCommMonoid P] [AddCommMonoid N] variable [Module R M] [Module R P] [Module R N] theorem isArtinian_of_injective (f : M →ₗ[R] P) (h : Function.Injective f) [IsArtinian R P] : @@ -86,10 +87,6 @@ theorem isArtinian_of_surjective (f : M →ₗ[R] P) (hf : Function.Surjective f show A.comap f < B.comap f from Submodule.comap_strictMono_of_surjective hf hAB) (InvImage.wf (Submodule.comap f) IsWellFounded.wf)⟩ -instance isArtinian_of_quotient_of_artinian - (N : Submodule R M) [IsArtinian R M] : IsArtinian R (M ⧸ N) := - isArtinian_of_surjective M (Submodule.mkQ N) (Submodule.Quotient.mk_surjective N) - variable {M} instance isArtinian_range (f : M →ₗ[R] P) [IsArtinian R M] : IsArtinian R (LinearMap.range f) := @@ -101,74 +98,17 @@ theorem isArtinian_of_linearEquiv (f : M ≃ₗ[R] P) [IsArtinian R M] : IsArtin theorem LinearEquiv.isArtinian_iff (f : M ≃ₗ[R] P) : IsArtinian R M ↔ IsArtinian R P := ⟨fun _ ↦ isArtinian_of_linearEquiv f, fun _ ↦ isArtinian_of_linearEquiv f.symm⟩ -theorem isArtinian_of_range_eq_ker [IsArtinian R M] [IsArtinian R P] (f : M →ₗ[R] N) (g : N →ₗ[R] P) - (h : LinearMap.range f = LinearMap.ker g) : IsArtinian R N := - wellFounded_lt_exact_sequence (LinearMap.range f) (Submodule.map (f.ker.liftQ f le_rfl)) - (Submodule.comap (f.ker.liftQ f le_rfl)) - (Submodule.comap g.rangeRestrict) (Submodule.map g.rangeRestrict) - (Submodule.gciMapComap <| LinearMap.ker_eq_bot.mp <| Submodule.ker_liftQ_eq_bot _ _ _ le_rfl) - (Submodule.giMapComap g.surjective_rangeRestrict) - (by simp [Submodule.map_comap_eq, inf_comm, Submodule.range_liftQ]) - (by simp [Submodule.comap_map_eq, h]) - -theorem isArtinian_iff_submodule_quotient (S : Submodule R P) : - IsArtinian R P ↔ IsArtinian R S ∧ IsArtinian R (P ⧸ S) := by - refine ⟨fun h ↦ ⟨inferInstance, inferInstance⟩, fun ⟨_, _⟩ ↦ ?_⟩ - apply isArtinian_of_range_eq_ker S.subtype S.mkQ - rw [Submodule.ker_mkQ, Submodule.range_subtype] - -instance isArtinian_prod [IsArtinian R M] [IsArtinian R P] : IsArtinian R (M × P) := - isArtinian_of_range_eq_ker (LinearMap.inl R M P) (LinearMap.snd R M P) (LinearMap.range_inl R M P) - instance (priority := 100) isArtinian_of_finite [Finite M] : IsArtinian R M := ⟨Finite.wellFounded_of_trans_of_irrefl _⟩ -- Porting note: elab_as_elim can only be global and cannot be changed on an imported decl -- attribute [local elab_as_elim] Finite.induction_empty_option -instance isArtinian_sup (M₁ M₂ : Submodule R P) [IsArtinian R M₁] [IsArtinian R M₂] : - IsArtinian R ↥(M₁ ⊔ M₂) := by - have := isArtinian_range (M₁.subtype.coprod M₂.subtype) - rwa [LinearMap.range_coprod, Submodule.range_subtype, Submodule.range_subtype] at this - -variable {ι : Type*} [Finite ι] - -instance isArtinian_pi : - ∀ {M : ι → Type*} [∀ i, AddCommGroup (M i)] - [∀ i, Module R (M i)] [∀ i, IsArtinian R (M i)], IsArtinian R (∀ i, M i) := by - apply Finite.induction_empty_option _ _ _ ι - · exact fun e h ↦ isArtinian_of_linearEquiv (LinearEquiv.piCongrLeft R _ e) - · infer_instance - · exact fun ih ↦ isArtinian_of_linearEquiv (LinearEquiv.piOptionEquivProd R).symm - -/-- A version of `isArtinian_pi` for non-dependent functions. We need this instance because -sometimes Lean fails to apply the dependent version in non-dependent settings (e.g., it fails to -prove that `ι → ℝ` is finite dimensional over `ℝ`). -/ -instance isArtinian_pi' [IsArtinian R M] : IsArtinian R (ι → M) := - isArtinian_pi - ---Porting note (https://github.com/leanprover-community/mathlib4/issues/10754): new instance -instance isArtinian_finsupp [IsArtinian R M] : IsArtinian R (ι →₀ M) := - isArtinian_of_linearEquiv (Finsupp.linearEquivFunOnFinite _ _ _).symm - -instance isArtinian_iSup : - ∀ {M : ι → Submodule R P} [∀ i, IsArtinian R (M i)], IsArtinian R ↥(⨆ i, M i) := by - apply Finite.induction_empty_option _ _ _ ι - · intro _ _ e h _ _; rw [← e.iSup_comp]; apply h - · intros; rw [iSup_of_empty]; infer_instance - · intro _ _ ih _ _; rw [iSup_option]; infer_instance - -end - -open Submodule Function - -section Ring - -variable {R M : Type*} [Ring R] [AddCommGroup M] [Module R M] +open Submodule theorem IsArtinian.finite_of_linearIndependent [Nontrivial R] [h : IsArtinian R M] {s : Set M} (hs : LinearIndependent R ((↑) : s → M)) : s.Finite := by - refine by_contradiction fun hf => (RelEmbedding.wellFounded_iff_no_descending_seq.1 h.wf).elim' ?_ + refine by_contradiction fun hf ↦ (RelEmbedding.wellFounded_iff_no_descending_seq.1 h.wf).elim' ?_ have f : ℕ ↪ s := Set.Infinite.natEmbedding s hf have : ∀ n, (↑) ∘ f '' { m | n ≤ m } ⊆ s := by rintro n x ⟨y, _, rfl⟩ @@ -179,10 +119,10 @@ theorem IsArtinian.finite_of_linearIndependent [Nontrivial R] [h : IsArtinian R rw [span_le_span_iff hs (this b) (this a), Set.image_subset_image_iff (Subtype.coe_injective.comp f.injective), Set.subset_def] simp only [Set.mem_setOf_eq] - exact ⟨fun hab x => le_trans hab, fun h => h _ le_rfl⟩ - exact ⟨⟨fun n => span R (Subtype.val ∘ f '' { m | n ≤ m }), fun x y => by + exact ⟨fun hab x ↦ hab.trans, (· _ le_rfl)⟩ + exact ⟨⟨fun n ↦ span R (Subtype.val ∘ f '' { m | n ≤ m }), fun x y ↦ by rw [le_antisymm_iff, ← this y x, ← this x y] - exact fun ⟨h₁, h₂⟩ => le_antisymm_iff.2 ⟨h₂, h₁⟩⟩, by + exact fun ⟨h₁, h₂⟩ ↦ le_antisymm_iff.2 ⟨h₂, h₁⟩⟩, by intro a b conv_rhs => rw [GT.gt, lt_iff_le_not_le, this, this, ← lt_iff_le_not_le] rfl⟩ @@ -219,12 +159,125 @@ theorem induction {P : Submodule R M → Prop} (hgt : ∀ I, (∀ J < I, P J) P I := WellFoundedLT.induction I hgt +open Function + +/-- Any injective endomorphism of an Artinian module is surjective. -/ +theorem surjective_of_injective_endomorphism (f : M →ₗ[R] M) (s : Injective f) : Surjective f := by + have h := ‹IsArtinian R M›; contrapose! h + rw [IsArtinian, WellFoundedLT, isWellFounded_iff] + refine (RelEmbedding.natGT (LinearMap.range <| f ^ ·) ?_).not_wellFounded_of_decreasing_seq + intro n + simp_rw [pow_succ, LinearMap.mul_eq_comp, LinearMap.range_comp, ← Submodule.map_top (f ^ n)] + refine Submodule.map_strictMono_of_injective (LinearMap.iterate_injective s n) (Ne.lt_top ?_) + rwa [Ne, LinearMap.range_eq_top] + +/-- Any injective endomorphism of an Artinian module is bijective. -/ +theorem bijective_of_injective_endomorphism (f : M →ₗ[R] M) (s : Injective f) : Bijective f := + ⟨s, surjective_of_injective_endomorphism f s⟩ + +/-- A sequence `f` of submodules of an artinian module, +with the supremum `f (n+1)` and the infimum of `f 0`, ..., `f n` being ⊤, +is eventually ⊤. -/ +theorem disjoint_partial_infs_eventually_top (f : ℕ → Submodule R M) + (h : ∀ n, Disjoint (partialSups (OrderDual.toDual ∘ f) n) (OrderDual.toDual (f (n + 1)))) : + ∃ n : ℕ, ∀ m, n ≤ m → f m = ⊤ := by + -- A little off-by-one cleanup first: + rsuffices ⟨n, w⟩ : ∃ n : ℕ, ∀ m, n ≤ m → OrderDual.toDual f (m + 1) = ⊤ + · use n + 1 + rintro (_ | m) p + · cases p + · apply w + exact Nat.succ_le_succ_iff.mp p + obtain ⟨n, w⟩ := monotone_stabilizes (partialSups (OrderDual.toDual ∘ f)) + refine ⟨n, fun m p => ?_⟩ + exact (h m).eq_bot_of_ge (sup_eq_left.1 <| (w (m + 1) <| le_add_right p).symm.trans <| w m p) + end IsArtinian namespace LinearMap variable [IsArtinian R M] +lemma eventually_iInf_range_pow_eq (f : Module.End R M) : + ∀ᶠ n in atTop, ⨅ m, LinearMap.range (f ^ m) = LinearMap.range (f ^ n) := by + obtain ⟨n, hn : ∀ m, n ≤ m → LinearMap.range (f ^ n) = LinearMap.range (f ^ m)⟩ := + IsArtinian.monotone_stabilizes f.iterateRange + refine eventually_atTop.mpr ⟨n, fun l hl ↦ le_antisymm (iInf_le _ _) (le_iInf fun m ↦ ?_)⟩ + rcases le_or_lt l m with h | h + · rw [← hn _ (hl.trans h), hn _ hl] + · exact f.iterateRange.monotone h.le + +end LinearMap + +end Semiring + +section Ring + +variable {R M P N : Type*} +variable [Ring R] [AddCommGroup M] [AddCommGroup P] [AddCommGroup N] +variable [Module R M] [Module R P] [Module R N] + +instance isArtinian_of_quotient_of_artinian + (N : Submodule R M) [IsArtinian R M] : IsArtinian R (M ⧸ N) := + isArtinian_of_surjective M (Submodule.mkQ N) (Submodule.Quotient.mk_surjective N) + +theorem isArtinian_of_range_eq_ker [IsArtinian R M] [IsArtinian R P] (f : M →ₗ[R] N) (g : N →ₗ[R] P) + (h : LinearMap.range f = LinearMap.ker g) : IsArtinian R N := + wellFounded_lt_exact_sequence (LinearMap.range f) (Submodule.map (f.ker.liftQ f le_rfl)) + (Submodule.comap (f.ker.liftQ f le_rfl)) + (Submodule.comap g.rangeRestrict) (Submodule.map g.rangeRestrict) + (Submodule.gciMapComap <| LinearMap.ker_eq_bot.mp <| Submodule.ker_liftQ_eq_bot _ _ _ le_rfl) + (Submodule.giMapComap g.surjective_rangeRestrict) + (by simp [Submodule.map_comap_eq, inf_comm, Submodule.range_liftQ]) + (by simp [Submodule.comap_map_eq, h]) + +theorem isArtinian_iff_submodule_quotient (S : Submodule R P) : + IsArtinian R P ↔ IsArtinian R S ∧ IsArtinian R (P ⧸ S) := by + refine ⟨fun h ↦ ⟨inferInstance, inferInstance⟩, fun ⟨_, _⟩ ↦ ?_⟩ + apply isArtinian_of_range_eq_ker S.subtype S.mkQ + rw [Submodule.ker_mkQ, Submodule.range_subtype] + +instance isArtinian_prod [IsArtinian R M] [IsArtinian R P] : IsArtinian R (M × P) := + isArtinian_of_range_eq_ker (LinearMap.inl R M P) (LinearMap.snd R M P) (LinearMap.range_inl R M P) + +instance isArtinian_sup (M₁ M₂ : Submodule R P) [IsArtinian R M₁] [IsArtinian R M₂] : + IsArtinian R ↥(M₁ ⊔ M₂) := by + have := isArtinian_range (M₁.subtype.coprod M₂.subtype) + rwa [LinearMap.range_coprod, Submodule.range_subtype, Submodule.range_subtype] at this + +variable {ι : Type*} [Finite ι] + +instance isArtinian_pi : + ∀ {M : ι → Type*} [Π i, AddCommGroup (M i)] + [Π i, Module R (M i)] [∀ i, IsArtinian R (M i)], IsArtinian R (Π i, M i) := by + apply Finite.induction_empty_option _ _ _ ι + · exact fun e h ↦ isArtinian_of_linearEquiv (LinearEquiv.piCongrLeft R _ e) + · infer_instance + · exact fun ih ↦ isArtinian_of_linearEquiv (LinearEquiv.piOptionEquivProd R).symm + +/-- A version of `isArtinian_pi` for non-dependent functions. We need this instance because +sometimes Lean fails to apply the dependent version in non-dependent settings (e.g., it fails to +prove that `ι → ℝ` is finite dimensional over `ℝ`). -/ +instance isArtinian_pi' [IsArtinian R M] : IsArtinian R (ι → M) := + isArtinian_pi + +--Porting note (https://github.com/leanprover-community/mathlib4/issues/10754): new instance +instance isArtinian_finsupp [IsArtinian R M] : IsArtinian R (ι →₀ M) := + isArtinian_of_linearEquiv (Finsupp.linearEquivFunOnFinite _ _ _).symm + +instance isArtinian_iSup : + ∀ {M : ι → Submodule R P} [∀ i, IsArtinian R (M i)], IsArtinian R ↥(⨆ i, M i) := by + apply Finite.induction_empty_option _ _ _ ι + · intro _ _ e h _ _; rw [← e.iSup_comp]; apply h + · intros; rw [iSup_of_empty]; infer_instance + · intro _ _ ih _ _; rw [iSup_option]; infer_instance + +open Submodule Function + +namespace LinearMap + +variable [IsArtinian R M] + /-- For any endomorphism of an Artinian module, any sufficiently high iterate has codisjoint kernel and range. -/ theorem eventually_codisjoint_ker_pow_range_pow (f : M →ₗ[R] M) : @@ -241,15 +294,6 @@ theorem eventually_codisjoint_ker_pow_range_pow (f : M →ₗ[R] M) : ← f.pow_apply m, ← mem_range (f := _), ← hn _ (n.le_add_left m), hn _ hm] exact LinearMap.mem_range_self (f ^ m) x -lemma eventually_iInf_range_pow_eq (f : Module.End R M) : - ∀ᶠ n in atTop, ⨅ m, LinearMap.range (f ^ m) = LinearMap.range (f ^ n) := by - obtain ⟨n, hn : ∀ m, n ≤ m → LinearMap.range (f ^ n) = LinearMap.range (f ^ m)⟩ := - IsArtinian.monotone_stabilizes f.iterateRange - refine eventually_atTop.mpr ⟨n, fun l hl ↦ le_antisymm (iInf_le _ _) (le_iInf fun m ↦ ?_)⟩ - rcases le_or_lt l m with h | h - · rw [← hn _ (hl.trans h), hn _ hl] - · exact f.iterateRange.monotone h.le - /-- This is the Fitting decomposition of the module `M` with respect to the endomorphism `f`. See also `LinearMap.isCompl_iSup_ker_pow_iInf_range_pow` for an alternative spelling. -/ @@ -271,46 +315,11 @@ theorem isCompl_iSup_ker_pow_iInf_range_pow [IsNoetherian R M] (f : M →ₗ[R] end LinearMap -namespace IsArtinian - -variable [IsArtinian R M] - -/-- Any injective endomorphism of an Artinian module is surjective. -/ -theorem surjective_of_injective_endomorphism (f : M →ₗ[R] M) (s : Injective f) : Surjective f := by - obtain ⟨n, hn⟩ := eventually_atTop.mp f.eventually_codisjoint_ker_pow_range_pow - specialize hn (n + 1) (n.le_add_right 1) - rw [codisjoint_iff, LinearMap.ker_eq_bot.mpr (LinearMap.iterate_injective s _), bot_sup_eq, - LinearMap.range_eq_top] at hn - exact LinearMap.surjective_of_iterate_surjective n.succ_ne_zero hn - -/-- Any injective endomorphism of an Artinian module is bijective. -/ -theorem bijective_of_injective_endomorphism (f : M →ₗ[R] M) (s : Injective f) : Bijective f := - ⟨s, surjective_of_injective_endomorphism f s⟩ - -/-- A sequence `f` of submodules of an artinian module, -with the supremum `f (n+1)` and the infimum of `f 0`, ..., `f n` being ⊤, -is eventually ⊤. -/ -theorem disjoint_partial_infs_eventually_top (f : ℕ → Submodule R M) - (h : ∀ n, Disjoint (partialSups (OrderDual.toDual ∘ f) n) (OrderDual.toDual (f (n + 1)))) : - ∃ n : ℕ, ∀ m, n ≤ m → f m = ⊤ := by - -- A little off-by-one cleanup first: - rsuffices ⟨n, w⟩ : ∃ n : ℕ, ∀ m, n ≤ m → OrderDual.toDual f (m + 1) = ⊤ - · use n + 1 - rintro (_ | m) p - · cases p - · apply w - exact Nat.succ_le_succ_iff.mp p - obtain ⟨n, w⟩ := monotone_stabilizes (partialSups (OrderDual.toDual ∘ f)) - refine ⟨n, fun m p => ?_⟩ - exact (h m).eq_bot_of_ge (sup_eq_left.1 <| (w (m + 1) <| le_add_right p).symm.trans <| w m p) - -end IsArtinian - end Ring -section CommRing +section CommSemiring -variable {R : Type*} (M : Type*) [CommRing R] [AddCommGroup M] [Module R M] [IsArtinian R M] +variable {R : Type*} (M : Type*) [CommSemiring R] [AddCommMonoid M] [Module R M] [IsArtinian R M] namespace IsArtinian @@ -334,13 +343,13 @@ theorem exists_pow_succ_smul_dvd (r : R) (x : M) : end IsArtinian -end CommRing +end CommSemiring -theorem isArtinian_of_submodule_of_artinian (R M) [Ring R] [AddCommGroup M] [Module R M] +theorem isArtinian_of_submodule_of_artinian (R M) [Semiring R] [AddCommMonoid M] [Module R M] (N : Submodule R M) (_ : IsArtinian R M) : IsArtinian R N := inferInstance /-- If `M / S / R` is a scalar tower, and `M / R` is Artinian, then `M / S` is also Artinian. -/ -theorem isArtinian_of_tower (R) {S M} [CommRing R] [Ring S] [AddCommGroup M] [Algebra R S] +theorem isArtinian_of_tower (R) {S M} [Semiring R] [Semiring S] [AddCommMonoid M] [SMul R S] [Module S M] [Module R M] [IsScalarTower R S M] (h : IsArtinian R M) : IsArtinian S M := ⟨(Submodule.restrictScalarsEmbedding R S M).wellFounded h.wf⟩ @@ -356,21 +365,26 @@ convenience in the commutative case. For a right Artinian ring, use `IsArtinian For equivalent definitions, see `Mathlib.RingTheory.Artinian.Ring`. -/ @[stacks 00J5] -abbrev IsArtinianRing (R) [Ring R] := +abbrev IsArtinianRing (R) [Semiring R] := IsArtinian R R -theorem isArtinianRing_iff {R} [Ring R] : IsArtinianRing R ↔ IsArtinian R R := Iff.rfl +theorem isArtinianRing_iff {R} [Semiring R] : IsArtinianRing R ↔ IsArtinian R R := Iff.rfl -instance DivisionRing.instIsArtinianRing {K : Type*} [DivisionRing K] : IsArtinianRing K := +instance DivisionSemiring.instIsArtinianRing {K : Type*} [DivisionSemiring K] : IsArtinianRing K := ⟨Finite.wellFounded_of_trans_of_irrefl _⟩ -theorem Ring.isArtinian_of_zero_eq_one {R} [Ring R] (h01 : (0 : R) = 1) : IsArtinianRing R := +instance DivisionRing.instIsArtinianRing {K : Type*} [DivisionRing K] : IsArtinianRing K := + inferInstance + +theorem Ring.isArtinian_of_zero_eq_one {R} [Semiring R] (h01 : (0 : R) = 1) : IsArtinianRing R := have := subsingleton_of_zero_eq_one h01 inferInstance instance (R) [CommRing R] [IsArtinianRing R] (I : Ideal R) : IsArtinianRing (R ⧸ I) := isArtinian_of_tower R inferInstance +open Submodule Function + theorem isArtinian_of_fg_of_artinian {R M} [Ring R] [AddCommGroup M] [Module R M] (N : Submodule R M) [IsArtinianRing R] (hN : N.FG) : IsArtinian R N := by let ⟨s, hs⟩ := hN @@ -390,7 +404,7 @@ theorem isArtinian_of_fg_of_artinian {R M} [Ring R] [AddCommGroup M] [Module R M instance isArtinian_of_fg_of_artinian' {R M} [Ring R] [AddCommGroup M] [Module R M] [IsArtinianRing R] [Module.Finite R M] : IsArtinian R M := - have : IsArtinian R (⊤ : Submodule R M) := isArtinian_of_fg_of_artinian _ Module.Finite.out + have : IsArtinian R (⊤ : Submodule R M) := isArtinian_of_fg_of_artinian _ Module.Finite.fg_top isArtinian_of_linearEquiv (LinearEquiv.ofTop (⊤ : Submodule R M) rfl) theorem IsArtinianRing.of_finite (R S) [CommRing R] [Ring S] [Algebra R S] @@ -403,18 +417,57 @@ theorem isArtinian_span_of_finite (R) {M} [Ring R] [AddCommGroup M] [Module R M] {A : Set M} (hA : A.Finite) : IsArtinian R (Submodule.span R A) := isArtinian_of_fg_of_artinian _ (Submodule.fg_def.mpr ⟨A, hA, rfl⟩) -theorem Function.Surjective.isArtinianRing {R} [Ring R] {S} [Ring S] {F} +theorem Function.Surjective.isArtinianRing {R} [Semiring R] {S} [Semiring S] {F} [FunLike F R S] [RingHomClass F R S] {f : F} (hf : Function.Surjective f) [H : IsArtinianRing R] : IsArtinianRing S := by rw [isArtinianRing_iff] at H ⊢ exact ⟨(Ideal.orderEmbeddingOfSurjective f hf).wellFounded H.wf⟩ +instance isArtinianRing_rangeS {R} [Semiring R] {S} [Semiring S] (f : R →+* S) [IsArtinianRing R] : + IsArtinianRing f.rangeS := + f.rangeSRestrict_surjective.isArtinianRing + instance isArtinianRing_range {R} [Ring R] {S} [Ring S] (f : R →+* S) [IsArtinianRing R] : IsArtinianRing f.range := - f.rangeRestrict_surjective.isArtinianRing + isArtinianRing_rangeS f + +theorem RingEquiv.isArtinianRing {R S} [Semiring R] [Semiring S] (f : R ≃+* S) + [IsArtinianRing R] : IsArtinianRing S := + f.surjective.isArtinianRing + +instance {R S} [Semiring R] [Semiring S] [IsArtinianRing R] [IsArtinianRing S] : + IsArtinianRing (R × S) := + Ideal.idealProdEquiv.toOrderEmbedding.wellFoundedLT + +instance {ι} [Finite ι] : ∀ {R : ι → Type*} [Π i, Semiring (R i)] [∀ i, IsArtinianRing (R i)], + IsArtinianRing (Π i, R i) := by + apply Finite.induction_empty_option _ _ _ ι + · exact fun e h ↦ RingEquiv.isArtinianRing (.piCongrLeft _ e) + · infer_instance + · exact fun ih ↦ RingEquiv.isArtinianRing (.symm .piOptionEquivProd) namespace IsArtinianRing +section CommSemiring + +variable (R : Type*) [CommSemiring R] [IsArtinianRing R] + +@[stacks 00J7] +lemma maximal_ideals_finite : {I : Ideal R | I.IsMaximal}.Finite := by + set Spec := {I : Ideal R | I.IsMaximal} + obtain ⟨_, ⟨s, rfl⟩, H⟩ := IsArtinian.set_has_minimal + (range (Finset.inf · Subtype.val : Finset Spec → Ideal R)) ⟨⊤, ∅, by simp⟩ + refine Set.finite_def.2 ⟨s, fun p ↦ ?_⟩ + classical + obtain ⟨q, hq1, hq2⟩ := p.2.isPrime.inf_le'.mp <| inf_eq_right.mp <| + inf_le_right.eq_of_not_lt (H (p ⊓ s.inf Subtype.val) ⟨insert p s, by simp⟩) + rwa [← Subtype.ext <| q.2.eq_of_le p.2.ne_top hq2] + +lemma subtype_isMaximal_finite : Finite {I : Ideal R | I.IsMaximal} := + (maximal_ideals_finite R).to_subtype + +end CommSemiring + variable {R : Type*} [CommRing R] [IsArtinianRing R] variable (R) in @@ -428,31 +481,20 @@ lemma isField_of_isDomain [IsDomain R] : IsField R := by rw [mul_eq_zero, sub_eq_zero] at hy exact ⟨_, hy.resolve_left <| pow_ne_zero _ hx⟩ +/- Does not hold in a commutative semiring: +consider {0, 0.5, 1} with ⊔ as + and ⊓ as *, then both {0} and {0, 0.5} are prime ideals. -/ instance isMaximal_of_isPrime (p : Ideal R) [p.IsPrime] : p.IsMaximal := Ideal.Quotient.maximal_of_isField _ (isField_of_isDomain _) lemma isPrime_iff_isMaximal (p : Ideal R) : p.IsPrime ↔ p.IsMaximal := ⟨fun _ ↦ isMaximal_of_isPrime p, fun h ↦ h.isPrime⟩ -variable (R) in -lemma primeSpectrum_finite : {I : Ideal R | I.IsPrime}.Finite := by - set Spec := {I : Ideal R | I.IsPrime} - obtain ⟨_, ⟨s, rfl⟩, H⟩ := IsArtinian.set_has_minimal - (range (Finset.inf · Subtype.val : Finset Spec → Ideal R)) ⟨⊤, ∅, by simp⟩ - refine Set.finite_def.2 ⟨s, fun p ↦ ?_⟩ - classical - obtain ⟨q, hq1, hq2⟩ := p.2.inf_le'.mp <| inf_eq_right.mp <| - inf_le_right.eq_of_not_lt (H (p ⊓ s.inf Subtype.val) ⟨insert p s, by simp⟩) - rwa [← Subtype.ext <| (@isMaximal_of_isPrime _ _ _ _ q.2).eq_of_le p.2.1 hq2] - variable (R) -@[stacks 00J7] -lemma maximal_ideals_finite : {I : Ideal R | I.IsMaximal}.Finite := by - simp_rw [← isPrime_iff_isMaximal] - apply primeSpectrum_finite R -@[local instance] lemma subtype_isMaximal_finite : Finite {I : Ideal R | I.IsMaximal} := - (maximal_ideals_finite R).to_subtype +lemma primeSpectrum_finite : {I : Ideal R | I.IsPrime}.Finite := by + simpa only [isPrime_iff_isMaximal] using maximal_ideals_finite R + +attribute [local instance] subtype_isMaximal_finite /-- A temporary field instance on the quotients by maximal ideals. -/ @[local instance] noncomputable def fieldOfSubtypeIsMaximal diff --git a/Mathlib/RingTheory/Artinian/Ring.lean b/Mathlib/RingTheory/Artinian/Ring.lean index 3c0e9a2fde380..a508c61431e11 100644 --- a/Mathlib/RingTheory/Artinian/Ring.lean +++ b/Mathlib/RingTheory/Artinian/Ring.lean @@ -85,7 +85,7 @@ theorem isNilpotent_jacobson_bot : IsNilpotent (Ideal.jacobson (⊥ : Ideal R)) section Localization -variable (S : Submonoid R) (L : Type*) [CommRing L] [Algebra R L] [IsLocalization S L] +variable (S : Submonoid R) (L : Type*) [CommSemiring L] [Algebra R L] [IsLocalization S L] include S /-- Localizing an artinian ring can only reduce the amount of elements. -/ diff --git a/Mathlib/RingTheory/Coprime/Lemmas.lean b/Mathlib/RingTheory/Coprime/Lemmas.lean index 0034b62252647..4ccc44de66126 100644 --- a/Mathlib/RingTheory/Coprime/Lemmas.lean +++ b/Mathlib/RingTheory/Coprime/Lemmas.lean @@ -22,6 +22,8 @@ lemmas about `Pow` since these are easiest to prove via `Finset.prod`. universe u v +open scoped Function -- required for scoped `on` notation + section IsCoprime variable {R : Type u} {I : Type v} [CommSemiring R] {x y z : R} {s : I → R} {t : Finset I} diff --git a/Mathlib/RingTheory/DedekindDomain/AdicValuation.lean b/Mathlib/RingTheory/DedekindDomain/AdicValuation.lean index 058f95c8bd4ad..77f7b9998fedb 100644 --- a/Mathlib/RingTheory/DedekindDomain/AdicValuation.lean +++ b/Mathlib/RingTheory/DedekindDomain/AdicValuation.lean @@ -364,7 +364,7 @@ variable {K} def adicValued : Valued K ℤₘ₀ := Valued.mk' v.valuation -theorem adicValued_apply {x : K} : (v.adicValued.v : _) x = v.valuation x := +theorem adicValued_apply {x : K} : v.adicValued.v x = v.valuation x := rfl variable (K) diff --git a/Mathlib/RingTheory/DedekindDomain/Different.lean b/Mathlib/RingTheory/DedekindDomain/Different.lean index 7b195346a4cf3..5ab55f608ea3d 100644 --- a/Mathlib/RingTheory/DedekindDomain/Different.lean +++ b/Mathlib/RingTheory/DedekindDomain/Different.lean @@ -462,7 +462,7 @@ open Submodule lemma differentialIdeal_le_fractionalIdeal_iff {I : FractionalIdeal B⁰ L} (hI : I ≠ 0) [NoZeroSMulDivisors A B] : - differentIdeal A B ≤ I ↔ (((I⁻¹ : _) : Submodule B L).restrictScalars A).map + differentIdeal A B ≤ I ↔ (((I⁻¹ :) : Submodule B L).restrictScalars A).map ((Algebra.trace K L).restrictScalars A) ≤ 1 := by rw [coeIdeal_differentIdeal A K L B, FractionalIdeal.inv_le_comm (by simp) hI, ← FractionalIdeal.coe_le_coe, FractionalIdeal.coe_dual_one] @@ -594,7 +594,7 @@ lemma pow_sub_one_dvd_differentIdeal_aux [IsFractionRing B L] [IsDedekindDomain rw [pow_add, hb, mul_assoc, mul_right_inj' (pow_ne_zero _ hPbot), pow_one, mul_comm] at ha exact ⟨_, ha.symm⟩ suffices ∀ x ∈ a, intTrace A B x ∈ p by - have hP : ((P ^ (e - 1) : _)⁻¹ : FractionalIdeal B⁰ L) = a / p.map (algebraMap A B) := by + have hP : ((P ^ (e - 1) :)⁻¹ : FractionalIdeal B⁰ L) = a / p.map (algebraMap A B) := by apply inv_involutive.injective simp only [inv_inv, ha, FractionalIdeal.coeIdeal_mul, inv_div, ne_eq, FractionalIdeal.coeIdeal_eq_zero, mul_div_assoc] diff --git a/Mathlib/RingTheory/Derivation/Basic.lean b/Mathlib/RingTheory/Derivation/Basic.lean index 5b01844f39b40..f530128d4efbe 100644 --- a/Mathlib/RingTheory/Derivation/Basic.lean +++ b/Mathlib/RingTheory/Derivation/Basic.lean @@ -137,7 +137,7 @@ theorem leibniz_pow (n : ℕ) : D (a ^ n) = n • a ^ (n - 1) • D a := by induction' n with n ihn · rw [pow_zero, map_one_eq_zero, zero_smul] · rcases (zero_le n).eq_or_lt with (rfl | hpos) - · erw [pow_one, one_smul, pow_zero, one_smul] + · simp · have : a * a ^ (n - 1) = a ^ n := by rw [← pow_succ', Nat.sub_add_cancel hpos] simp only [pow_succ', leibniz, ihn, smul_comm a n (_ : M), smul_smul a, add_smul, this, Nat.succ_eq_add_one, Nat.add_succ_sub_one, add_zero, one_nsmul] diff --git a/Mathlib/RingTheory/Derivation/ToSquareZero.lean b/Mathlib/RingTheory/Derivation/ToSquareZero.lean index e99816f3014ce..dfaa5966fbf27 100644 --- a/Mathlib/RingTheory/Derivation/ToSquareZero.lean +++ b/Mathlib/RingTheory/Derivation/ToSquareZero.lean @@ -117,7 +117,7 @@ lifts `A →ₐ[R] B` of the canonical map `A →ₐ[R] B ⧸ I`. -/ def derivationToSquareZeroEquivLift [IsScalarTower R A B] : Derivation R A I ≃ { f : A →ₐ[R] B // (Ideal.Quotient.mkₐ R I).comp f = IsScalarTower.toAlgHom R A (B ⧸ I) } := by refine ⟨fun d => ⟨liftOfDerivationToSquareZero I hI d, ?_⟩, fun f => - (derivationToSquareZeroOfLift I hI f.1 f.2 : _), ?_, ?_⟩ + (derivationToSquareZeroOfLift I hI f.1 f.2 :), ?_, ?_⟩ · ext x; exact liftOfDerivationToSquareZero_mk_apply I hI d x · intro d; ext x; exact add_sub_cancel_right (d x : B) (algebraMap A B x) · rintro ⟨f, hf⟩; ext x; exact sub_add_cancel (f x) (algebraMap A B x) diff --git a/Mathlib/RingTheory/Etale/Field.lean b/Mathlib/RingTheory/Etale/Field.lean index 3e6b3bd1c7814..be84c58cf9418 100644 --- a/Mathlib/RingTheory/Etale/Field.lean +++ b/Mathlib/RingTheory/Etale/Field.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 Andrew Yang. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang -/ -import Mathlib.RingTheory.Etale.Basic +import Mathlib.RingTheory.Etale.Pi import Mathlib.RingTheory.Unramified.Field /-! @@ -14,9 +14,14 @@ import Mathlib.RingTheory.Unramified.Field Let `K` be a field, `A` be a `K`-algebra and `L` be a field extension of `K`. - `Algebra.FormallyEtale.of_isSeparable`: - If `L` is separable over `K`, then `L` is formally etale over `K`. + If `L` is separable over `K`, then `L` is formally étale over `K`. - `Algebra.FormallyEtale.iff_isSeparable`: - If `L` is (essentially) of finite type over `K`, then `L/K` is etale iff `L/K` is separable. + If `L` is (essentially) of finite type over `K`, then `L/K` is étale iff `L/K` is separable. +- `Algebra.FormallyEtale.iff_exists_algEquiv_prod`: + If `A` is (essentially) of finite type over `K`, + then `A/K` is étale iff `A` is a finite product of separable field extensions. +- `Algebra.Etale.iff_exists_algEquiv_prod`: + `A/K` is étale iff `A` is a finite product of finite separable field extensions. ## References @@ -27,7 +32,7 @@ Let `K` be a field, `A` be a `K`-algebra and `L` be a field extension of `K`. universe u -variable (K L : Type u) [Field K] [Field L] [Algebra K L] +variable (K L A : Type u) [Field K] [Field L] [CommRing A] [Algebra K L] [Algebra K A] open Algebra Polynomial @@ -148,4 +153,61 @@ theorem iff_isSeparable [EssFiniteType K L] : FormallyEtale K L ↔ Algebra.IsSeparable K L := ⟨fun _ ↦ FormallyUnramified.isSeparable K L, fun _ ↦ of_isSeparable K L⟩ +attribute [local instance] + IsArtinianRing.subtype_isMaximal_finite IsArtinianRing.fieldOfSubtypeIsMaximal in +/-- +If `A` is an essentially of finite type algebra over a field `K`, then `A` is formally étale +over `K` if and only if `A` is a finite product of separable field extensions. +-/ +theorem iff_exists_algEquiv_prod [EssFiniteType K A] : + FormallyEtale K A ↔ + ∃ (I : Type u) (_ : Finite I) (Ai : I → Type u) (_ : ∀ i, Field (Ai i)) + (_ : ∀ i, Algebra K (Ai i)) (_ : A ≃ₐ[K] Π i, Ai i), + ∀ i, Algebra.IsSeparable K (Ai i) := by + classical + constructor + · intro H + have := FormallyUnramified.finite_of_free K A + have := FormallyUnramified.isReduced_of_field K A + have : IsArtinianRing A := isArtinian_of_tower K inferInstance + letI : Fintype {I : Ideal A | I.IsMaximal} := (nonempty_fintype _).some + let v (i : {I : Ideal A | I.IsMaximal}) : A := (IsArtinianRing.equivPi A).symm (Pi.single i 1) + let e : A ≃ₐ[K] _ := { __ := IsArtinianRing.equivPi A, commutes' := fun r ↦ rfl } + have := (FormallyEtale.iff_of_equiv e).mp inferInstance + rw [FormallyEtale.pi_iff] at this + exact ⟨_, inferInstance, _, _, _, e, fun I ↦ (iff_isSeparable _ _).mp inferInstance⟩ + · intro ⟨I, _, Ai, _, _, e, _⟩ + rw [FormallyEtale.iff_of_equiv e, FormallyEtale.pi_iff] + have (i) : EssFiniteType K (Ai i) := by + letI := ((Pi.evalRingHom Ai i).comp e.toRingHom).toAlgebra + have : IsScalarTower K A (Ai i) := + .of_algebraMap_eq fun r ↦ by simp [RingHom.algebraMap_toAlgebra] + have : Algebra.FiniteType A (Ai i) := .of_surjective inferInstance (Algebra.ofId _ _) + (RingHomSurjective.is_surjective (σ := Pi.evalRingHom Ai i).comp e.surjective) + exact EssFiniteType.comp K A (Ai i) + exact fun I ↦ (iff_isSeparable _ _).mpr inferInstance + end Algebra.FormallyEtale + +attribute [local instance] + IsArtinianRing.subtype_isMaximal_finite IsArtinianRing.fieldOfSubtypeIsMaximal in +/-- +`A` is étale over a field `K` if and only if +`A` is a finite product of finite separable field extensions. +-/ +theorem Algebra.Etale.iff_exists_algEquiv_prod : + Etale K A ↔ + ∃ (I : Type u) (_ : Finite I) (Ai : I → Type u) (_ : ∀ i, Field (Ai i)) + (_ : ∀ i, Algebra K (Ai i)) (_ : A ≃ₐ[K] Π i, Ai i), + ∀ i, Module.Finite K (Ai i) ∧ Algebra.IsSeparable K (Ai i) := by + constructor + · intro H + obtain ⟨I, _, Ai, _, _, e, _⟩ := (FormallyEtale.iff_exists_algEquiv_prod K A).mp inferInstance + have := FormallyUnramified.finite_of_free K A + exact ⟨_, ‹_›, _, _, _, e, fun i ↦ ⟨.of_surjective ((LinearMap.proj i).comp e.toLinearMap) + ((Function.surjective_eval i).comp e.surjective), inferInstance⟩⟩ + · intro ⟨I, _, Ai, _, _, e, H⟩ + choose h₁ h₂ using H + have := Module.Finite.of_surjective e.symm.toLinearMap e.symm.surjective + refine ⟨?_, FinitePresentation.of_finiteType.mp inferInstance⟩ + exact (FormallyEtale.iff_exists_algEquiv_prod K A).mpr ⟨_, inferInstance, _, _, _, e, h₂⟩ diff --git a/Mathlib/RingTheory/Etale/Kaehler.lean b/Mathlib/RingTheory/Etale/Kaehler.lean new file mode 100644 index 0000000000000..509af402e838f --- /dev/null +++ b/Mathlib/RingTheory/Etale/Kaehler.lean @@ -0,0 +1,53 @@ +/- +Copyright (c) 2024 Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Andrew Yang +-/ +import Mathlib.RingTheory.Etale.Basic +import Mathlib.RingTheory.Kaehler.JacobiZariski +import Mathlib.RingTheory.Localization.BaseChange +import Mathlib.RingTheory.Smooth.Kaehler + +/-! +# The differential module and etale algebras + +## Main results +`KaehlerDifferential.tensorKaehlerEquivOfFormallyEtale`: +The canonical isomorphism `T ⊗[S] Ω[S⁄R] ≃ₗ[T] Ω[T⁄R]` for `T` a formally etale `S`-algebra. +-/ + +universe u + +variable (R S T : Type u) [CommRing R] [CommRing S] [CommRing T] +variable [Algebra R S] [Algebra R T] [Algebra S T] [IsScalarTower R S T] + +open TensorProduct + +/-- +The canonical isomorphism `T ⊗[S] Ω[S⁄R] ≃ₗ[T] Ω[T⁄R]` for `T` a formally etale `S`-algebra. +Also see `S ⊗[R] Ω[A⁄R] ≃ₗ[S] Ω[S ⊗[R] A⁄S]` at `KaehlerDifferential.tensorKaehlerEquiv`. +-/ +@[simps! apply] noncomputable +def KaehlerDifferential.tensorKaehlerEquivOfFormallyEtale [Algebra.FormallyEtale S T] : + T ⊗[S] Ω[S⁄R] ≃ₗ[T] Ω[T⁄R] := by + refine LinearEquiv.ofBijective (mapBaseChange R S T) + ⟨?_, fun x ↦ (KaehlerDifferential.exact_mapBaseChange_map R S T x).mp (Subsingleton.elim _ _)⟩ + rw [injective_iff_map_eq_zero] + intros x hx + obtain ⟨x, rfl⟩ := (Algebra.H1Cotangent.exact_δ_mapBaseChange R S T x).mp hx + rw [Subsingleton.elim x 0, map_zero] + +lemma KaehlerDifferential.isBaseChange_of_formallyEtale [Algebra.FormallyEtale S T] : + IsBaseChange T (map R R S T) := by + show Function.Bijective _ + convert (tensorKaehlerEquivOfFormallyEtale R S T).bijective using 1 + show _ = ((tensorKaehlerEquivOfFormallyEtale + R S T).toLinearMap.restrictScalars S : T ⊗[S] Ω[S⁄R] → _) + congr! + ext + simp + +instance KaehlerDifferential.isLocalizedModule_map (M : Submonoid S) [IsLocalization M T] : + IsLocalizedModule M (map R R S T) := + have := Algebra.FormallyEtale.of_isLocalization (Rₘ := T) M + (isLocalizedModule_iff_isBaseChange M T _).mpr (isBaseChange_of_formallyEtale R S T) diff --git a/Mathlib/RingTheory/Etale/Pi.lean b/Mathlib/RingTheory/Etale/Pi.lean new file mode 100644 index 0000000000000..0e50eccd56627 --- /dev/null +++ b/Mathlib/RingTheory/Etale/Pi.lean @@ -0,0 +1,36 @@ +/- +Copyright (c) 2024 Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Andrew Yang +-/ +import Mathlib.RingTheory.Smooth.Pi +import Mathlib.RingTheory.Unramified.Pi +import Mathlib.RingTheory.Etale.Basic + +/-! + +# Formal-étaleness of finite products of rings + +## Main result + +- `Algebra.FormallyEtale.pi_iff`: If `I` is finite, `Π i : I, A i` is `R`-formally-étale + if and only if each `A i` is `R`-formally-étale. + +-/ + +universe u v + +namespace Algebra.FormallyEtale + +variable {R : Type (max u v)} {I : Type u} (A : I → Type (max u v)) +variable [CommRing R] [∀ i, CommRing (A i)] [∀ i, Algebra R (A i)] + +theorem pi_iff [Finite I] : + FormallyEtale R (Π i, A i) ↔ ∀ i, FormallyEtale R (A i) := by + simp_rw [FormallyEtale.iff_unramified_and_smooth, forall_and] + rw [FormallyUnramified.pi_iff A, FormallySmooth.pi_iff A] + +instance [Finite I] [∀ i, FormallyEtale R (A i)] : FormallyEtale R (Π i, A i) := + .of_unramified_and_smooth + +end Algebra.FormallyEtale diff --git a/Mathlib/RingTheory/Finiteness/Basic.lean b/Mathlib/RingTheory/Finiteness/Basic.lean index 73248997ca909..bb1058f07fbc1 100644 --- a/Mathlib/RingTheory/Finiteness/Basic.lean +++ b/Mathlib/RingTheory/Finiteness/Basic.lean @@ -251,13 +251,13 @@ theorem equiv_iff (e : M ≃ₗ[R] N) : Module.Finite R M ↔ Module.Finite R N instance ulift [Module.Finite R M] : Module.Finite R (ULift M) := equiv ULift.moduleEquiv.symm -theorem iff_fg {N : Submodule R M} : Module.Finite R N ↔ N.FG := Module.finite_def.trans (fg_top _) +theorem iff_fg {N : Submodule R M} : Module.Finite R N ↔ N.FG := Module.finite_def.trans N.fg_top variable (R M) instance bot : Module.Finite R (⊥ : Submodule R M) := iff_fg.mpr fg_bot -instance top [Module.Finite R M] : Module.Finite R (⊤ : Submodule R M) := iff_fg.mpr out +instance top [Module.Finite R M] : Module.Finite R (⊤ : Submodule R M) := iff_fg.mpr fg_top variable {M} diff --git a/Mathlib/RingTheory/Finiteness/Defs.lean b/Mathlib/RingTheory/Finiteness/Defs.lean index fd8c45bb0be8d..9c52d2ce2bd80 100644 --- a/Mathlib/RingTheory/Finiteness/Defs.lean +++ b/Mathlib/RingTheory/Finiteness/Defs.lean @@ -103,9 +103,9 @@ variable (R A B M N : Type*) /-- A module over a semiring is `Module.Finite` if it is finitely generated as a module. -/ protected class Module.Finite [Semiring R] [AddCommMonoid M] [Module R M] : Prop where - out : (⊤ : Submodule R M).FG + fg_top : (⊤ : Submodule R M).FG -attribute [inherit_doc Module.Finite] Module.Finite.out +attribute [inherit_doc Module.Finite] Module.Finite.fg_top namespace Module @@ -131,7 +131,7 @@ variable {R M N} /-- See also `Module.Finite.exists_fin'`. -/ lemma exists_fin [Module.Finite R M] : ∃ (n : ℕ) (s : Fin n → M), Submodule.span R (range s) = ⊤ := - Submodule.fg_iff_exists_fin_generating_family.mp out + Submodule.fg_iff_exists_fin_generating_family.mp fg_top end Finite diff --git a/Mathlib/RingTheory/Finiteness/Nilpotent.lean b/Mathlib/RingTheory/Finiteness/Nilpotent.lean index 17b91ef86e77f..b2e0915fb0405 100644 --- a/Mathlib/RingTheory/Finiteness/Nilpotent.lean +++ b/Mathlib/RingTheory/Finiteness/Nilpotent.lean @@ -20,7 +20,7 @@ namespace Finite theorem Module.End.isNilpotent_iff_of_finite [Module.Finite R M] {f : End R M} : IsNilpotent f ↔ ∀ m : M, ∃ n : ℕ, (f ^ n) m = 0 := by refine ⟨fun ⟨n, hn⟩ m ↦ ⟨n, by simp [hn]⟩, fun h ↦ ?_⟩ - rcases Module.Finite.out (R := R) (M := M) with ⟨S, hS⟩ + rcases Module.Finite.fg_top (R := R) (M := M) with ⟨S, hS⟩ choose g hg using h use Finset.sup S g ext m diff --git a/Mathlib/RingTheory/Finiteness/TensorProduct.lean b/Mathlib/RingTheory/Finiteness/TensorProduct.lean index e1a686b2c3209..abfabac2a55e3 100644 --- a/Mathlib/RingTheory/Finiteness/TensorProduct.lean +++ b/Mathlib/RingTheory/Finiteness/TensorProduct.lean @@ -85,7 +85,7 @@ noncomputable local instance instance Module.Finite.base_change [CommSemiring R] [Semiring A] [Algebra R A] [AddCommMonoid M] [Module R M] [h : Module.Finite R M] : Module.Finite A (TensorProduct R A M) := by classical - obtain ⟨s, hs⟩ := h.out + obtain ⟨s, hs⟩ := h.fg_top refine ⟨⟨s.image (TensorProduct.mk R A M 1), eq_top_iff.mpr ?_⟩⟩ rintro x - induction x with @@ -103,7 +103,7 @@ instance Module.Finite.base_change [CommSemiring R] [Semiring A] [Algebra R A] [ instance Module.Finite.tensorProduct [CommSemiring R] [AddCommMonoid M] [Module R M] [AddCommMonoid N] [Module R N] [hM : Module.Finite R M] [hN : Module.Finite R N] : Module.Finite R (TensorProduct R M N) where - out := (TensorProduct.map₂_mk_top_top_eq_top R M N).subst (hM.out.map₂ _ hN.out) + fg_top := (TensorProduct.map₂_mk_top_top_eq_top R M N).subst (hM.fg_top.map₂ _ hN.fg_top) end ModuleAndAlgebra diff --git a/Mathlib/RingTheory/Flat/Basic.lean b/Mathlib/RingTheory/Flat/Basic.lean index 66b8a73a95d3f..aa58cec320090 100644 --- a/Mathlib/RingTheory/Flat/Basic.lean +++ b/Mathlib/RingTheory/Flat/Basic.lean @@ -107,9 +107,6 @@ theorem iff_rTensor_injective' : rewrite [← rTensor_comp_apply] at hx₀ rw [(injective_iff_map_eq_zero _).mp (h hfg) y hx₀, LinearMap.map_zero] -@[deprecated (since := "2024-03-29")] -alias lTensor_inj_iff_rTensor_inj := LinearMap.lTensor_inj_iff_rTensor_inj - /-- The `lTensor`-variant of `iff_rTensor_injective`. . -/ theorem iff_lTensor_injective : Module.Flat R M ↔ ∀ ⦃I : Ideal R⦄ (_ : I.FG), Function.Injective (lTensor M I.subtype) := by @@ -246,7 +243,8 @@ theorem iff_characterModule_baer : Flat R M ↔ Module.Baer R (CharacterModule M simp_rw [iff_rTensor_injective', Baer, rTensor_injective_iff_lcomp_surjective, Surjective, DFunLike.ext_iff, Subtype.forall]; rfl -/-- `CharacterModule M` is an injective module iff `M` is flat. -/ +/-- `CharacterModule M` is an injective module iff `M` is flat. +See [Lambek_1964] for a self-contained proof. -/ theorem iff_characterModule_injective [Small.{v} R] : Flat R M ↔ Module.Injective R (CharacterModule M) := iff_characterModule_baer.trans Module.Baer.iff_injective @@ -259,9 +257,6 @@ theorem rTensor_preserves_injective_linearMap {N' : Type*} [AddCommGroup N'] [Mo Function.Injective (L.rTensor M) := rTensor_injective_iff_lcomp_surjective.2 ((iff_characterModule_baer.1 h).extension_property _ hL) -@[deprecated (since := "2024-03-29")] -alias preserves_injective_linearMap := rTensor_preserves_injective_linearMap - instance {S} [CommRing S] [Algebra R S] [Module S M] [IsScalarTower R S M] [Flat S M] [Flat R N] : Flat S (M ⊗[R] N) := (iff_rTensor_injective' _ _).mpr fun I ↦ by diff --git a/Mathlib/RingTheory/Flat/EquationalCriterion.lean b/Mathlib/RingTheory/Flat/EquationalCriterion.lean index 763d9dc2c83db..ccb3ba16c850c 100644 --- a/Mathlib/RingTheory/Flat/EquationalCriterion.lean +++ b/Mathlib/RingTheory/Flat/EquationalCriterion.lean @@ -25,9 +25,9 @@ The equational criterion for flatness can be stated in the following form (`Module.Flat.iff_forall_exists_factorization`). Let $M$ be an $R$-module. Then the following two conditions are equivalent: * $M$ is flat. -* For all finite index types $\iota$, all elements $f \in R^{\iota}$, and all homomorphisms +* For all finite index types $\iota$, all elements $f \in R^{\iota}$, and all linear maps $x \colon R^{\iota} \to M$ such that $x(f) = 0$, there exist a finite index type $\kappa$ and module -homomorphisms $a \colon R^{\iota} \to R^{\kappa}$ and $y \colon R^{\kappa} \to M$ such +linear maps $a \colon R^{\iota} \to R^{\kappa}$ and $y \colon R^{\kappa} \to M$ such that $x = y \circ a$ and $a(f) = 0$. Of course, the module $R^\iota$ in this statement can be replaced by an arbitrary free module @@ -36,12 +36,12 @@ Of course, the module $R^\iota$ in this statement can be replaced by an arbitrar We also have the following strengthening of the equational criterion for flatness (`Module.Flat.exists_factorization_of_comp_eq_zero_of_free`): Let $M$ be a flat module. Let $K$ and $N$ be finite $R$-modules with $N$ free, and let $f \colon K \to N$ and -$x \colon N \to M$ be homomorphisms such that $x \circ f = 0$. Then there exist a finite index type -$\kappa$ and module homomorphisms $a \colon N \to R^{\kappa}$ and $y \colon R^{\kappa} \to M$ such +$x \colon N \to M$ be linear maps such that $x \circ f = 0$. Then there exist a finite index type +$\kappa$ and $R$-linear maps $a \colon N \to R^{\kappa}$ and $y \colon R^{\kappa} \to M$ such that $x = y \circ a$ and $a \circ f = 0$. We recover the usual equational criterion for flatness if $K = R$ and $N = R^\iota$. This is used in the proof of Lazard's theorem. -We conclude that every homomorphism from a finitely presented module to a flat module factors +We conclude that every linear map from a finitely presented module to a flat module factors through a finite free module (`Module.Flat.exists_factorization_of_isFinitelyPresented`). ## References @@ -103,12 +103,12 @@ Let $M$ be a module over a commutative ring $R$. The following are equivalent: * For all ideals $I \subseteq R$, the map $I \otimes M \to M$ is injective. * Every $\sum_i f_i \otimes x_i$ that vanishes in $R \otimes M$ vanishes trivially. * Every relation $\sum_i f_i x_i = 0$ in $M$ is trivial. -* For all finite index types $\iota$, all elements $f \in R^{\iota}$, and all homomorphisms +* For all finite index types $\iota$, all elements $f \in R^{\iota}$, and all linear maps $x \colon R^{\iota} \to M$ such that $x(f) = 0$, there exist a finite index type $\kappa$ and module -homomorphisms $a \colon R^{\iota} \to R^{\kappa}$ and $y \colon R^{\kappa} \to M$ such +linear maps $a \colon R^{\iota} \to R^{\kappa}$ and $y \colon R^{\kappa} \to M$ such that $x = y \circ a$ and $a(f) = 0$. -* For all finite free modules $N$, all elements $f \in N$, and all homomorphisms $x \colon N \to M$ -such that $x(f) = 0$, there exist a finite index type $\kappa$ and module homomorphisms +* For all finite free modules $N$, all elements $f \in N$, and all linear maps $x \colon N \to M$ +such that $x(f) = 0$, there exist a finite index type $\kappa$ and $R$-linear maps $a \colon N \to R^{\kappa}$ and $y \colon R^{\kappa} \to M$ such that $x = y \circ a$ and $a(f) = 0$. -/ theorem tfae_equational_criterion : List.TFAE [ @@ -203,8 +203,8 @@ theorem of_forall_isTrivialRelation (hfx : ∀ {ι : Type u} [Fintype ι] {f : [Stacks 00HK](https://stacks.math.columbia.edu/tag/00HK), alternate form. A module $M$ is flat if and only if for all finite free modules $R^\iota$, -all $f \in R^{\iota}$, and all homomorphisms $x \colon R^{\iota} \to M$ such that $x(f) = 0$, there -exist a finite free module $R^\kappa$ and homomorphisms $a \colon R^{\iota} \to R^{\kappa}$ and +all $f \in R^{\iota}$, and all linear maps $x \colon R^{\iota} \to M$ such that $x(f) = 0$, there +exist a finite free module $R^\kappa$ and linear maps $a \colon R^{\iota} \to R^{\kappa}$ and $y \colon R^{\kappa} \to M$ such that $x = y \circ a$ and $a(f) = 0$. -/ theorem iff_forall_exists_factorization : Flat R M ↔ ∀ {ι : Type u} [Fintype ι] {f : ι →₀ R} {x : (ι →₀ R) →ₗ[R] M}, x f = 0 → @@ -215,8 +215,8 @@ theorem iff_forall_exists_factorization : Flat R M ↔ [Stacks 00HK](https://stacks.math.columbia.edu/tag/00HK), forward direction, alternate form. Let $M$ be a flat module. Let $R^\iota$ be a finite free module, let $f \in R^{\iota}$ be an -element, and let $x \colon R^{\iota} \to M$ be a homomorphism such that $x(f) = 0$. Then there -exist a finite free module $R^\kappa$ and homomorphisms $a \colon R^{\iota} \to R^{\kappa}$ and +element, and let $x \colon R^{\iota} \to M$ be a linear map such that $x(f) = 0$. Then there +exist a finite free module $R^\kappa$ and linear maps $a \colon R^{\iota} \to R^{\kappa}$ and $y \colon R^{\kappa} \to M$ such that $x = y \circ a$ and $a(f) = 0$. -/ theorem exists_factorization_of_apply_eq_zero [Flat R M] {ι : Type u} [_root_.Finite ι] {f : ι →₀ R} {x : (ι →₀ R) →ₗ[R] M} (h : x f = 0) : @@ -228,8 +228,8 @@ theorem exists_factorization_of_apply_eq_zero [Flat R M] {ι : Type u} [_root_.F [Stacks 00HK](https://stacks.math.columbia.edu/tag/00HK), backward direction, alternate form. Let $M$ be a module over a commutative ring $R$. Suppose that for all finite free modules $R^\iota$, -all $f \in R^{\iota}$, and all homomorphisms $x \colon R^{\iota} \to M$ such that $x(f) = 0$, there -exist a finite free module $R^\kappa$ and homomorphisms $a \colon R^{\iota} \to R^{\kappa}$ and +all $f \in R^{\iota}$, and all linear maps $x \colon R^{\iota} \to M$ such that $x(f) = 0$, there +exist a finite free module $R^\kappa$ and linear maps $a \colon R^{\iota} \to R^{\kappa}$ and $y \colon R^{\kappa} \to M$ such that $x = y \circ a$ and $a(f) = 0$. Then $M$ is flat. -/ theorem of_forall_exists_factorization (h : ∀ {ι : Type u} [Fintype ι] {f : ι →₀ R} {x : (ι →₀ R) →ₗ[R] M}, x f = 0 → @@ -240,8 +240,8 @@ theorem of_forall_exists_factorization (h : ∀ {ι : Type u} [Fintype ι] {f : forward direction, second alternate form. Let $M$ be a flat module over a commutative ring $R$. Let $N$ be a finite free module over $R$, -let $f \in N$, and let $x \colon N \to M$ be a homomorphism such that $x(f) = 0$. Then there exist a -finite index type $\kappa$ and module homomorphisms $a \colon N \to R^{\kappa}$ and +let $f \in N$, and let $x \colon N \to M$ be a linear map such that $x(f) = 0$. Then there exist a +finite index type $\kappa$ and $R$-linear maps $a \colon N \to R^{\kappa}$ and $y \colon R^{\kappa} \to M$ such that $x = y \circ a$ and $a(f) = 0$. -/ theorem exists_factorization_of_apply_eq_zero_of_free [Flat R M] {N : Type u} [AddCommGroup N] [Module R N] [Free R N] [Module.Finite R N] {f : N} {x : N →ₗ[R] M} (h : x f = 0) : @@ -250,8 +250,8 @@ theorem exists_factorization_of_apply_eq_zero_of_free [Flat R M] {N : Type u} [A exact ((tfae_equational_criterion R M).out 0 5 rfl rfl).mp ‹Flat R M› h /-- Let $M$ be a flat module. Let $K$ and $N$ be finite $R$-modules with $N$ -free, and let $f \colon K \to N$ and $x \colon N \to M$ be homomorphisms such that -$x \circ f = 0$. Then there exist a finite index type $\kappa$ and module homomorphisms +free, and let $f \colon K \to N$ and $x \colon N \to M$ be linear maps such that +$x \circ f = 0$. Then there exist a finite index type $\kappa$ and $R$-linear maps $a \colon N \to R^{\kappa}$ and $y \colon R^{\kappa} \to M$ such that $x = y \circ a$ and $a \circ f = 0$. -/ theorem exists_factorization_of_comp_eq_zero_of_free [Flat R M] {K N : Type u} [AddCommGroup K] @@ -273,10 +273,10 @@ theorem exists_factorization_of_comp_eq_zero_of_free [Flat R M] {K N : Type u} [ use κ₂, hκ₂, a₂ ∘ₗ a₁, y₂ simp_rw [comp_assoc] exact ⟨trivial, sup_le (ha₁.trans (ker_le_ker_comp _ _)) ha₂⟩ - convert this ⊤ Finite.out + convert this ⊤ Finite.fg_top simp only [top_le_iff, ker_eq_top] -/-- Every homomorphism from a finitely presented module to a flat module factors through a finite +/-- Every linear map from a finitely presented module to a flat module factors through a finite free module. -/ theorem exists_factorization_of_isFinitelyPresented [Flat R M] {P : Type u} [AddCommGroup P] [Module R P] [FinitePresentation R P] (h₁ : P →ₗ[R] M) : diff --git a/Mathlib/RingTheory/Flat/FaithfullyFlat.lean b/Mathlib/RingTheory/Flat/FaithfullyFlat.lean index 39e1cd3c439b7..600a25333f1bd 100644 --- a/Mathlib/RingTheory/Flat/FaithfullyFlat.lean +++ b/Mathlib/RingTheory/Flat/FaithfullyFlat.lean @@ -55,7 +55,7 @@ variable (R : Type u) (M : Type v) [CommRing R] [AddCommGroup M] [Module R M] /-- A module `M` over a commutative ring `R` is *faithfully flat* if it is flat and, -for all `R`-module homomorphism `f : N → N'` such that `id ⊗ f = 0`, we have `f = 0`. +for all `R`-linear maps `f : N → N'` such that `id ⊗ f = 0`, we have `f = 0`. -/ @[mk_iff] class FaithfullyFlat extends Module.Flat R M : Prop where submodule_ne_top : ∀ ⦃m : Ideal R⦄ (_ : Ideal.IsMaximal m), m • (⊤ : Submodule R M) ≠ ⊤ diff --git a/Mathlib/RingTheory/HahnSeries/Addition.lean b/Mathlib/RingTheory/HahnSeries/Addition.lean index 817b97d65f6d7..01acbffb87646 100644 --- a/Mathlib/RingTheory/HahnSeries/Addition.lean +++ b/Mathlib/RingTheory/HahnSeries/Addition.lean @@ -83,10 +83,9 @@ and the additive opposite of Hahn series over `Γ` with coefficients `R`. def addOppositeEquiv : HahnSeries Γ (Rᵃᵒᵖ) ≃+ (HahnSeries Γ R)ᵃᵒᵖ where toFun x := .op ⟨fun a ↦ (x.coeff a).unop, by convert x.isPWO_support; ext; simp⟩ invFun x := ⟨fun a ↦ .op (x.unop.coeff a), by convert x.unop.isPWO_support; ext; simp⟩ - left_inv x := by ext; simp + left_inv x := by simp right_inv x := by apply AddOpposite.unop_injective - ext simp map_add' x y := by apply AddOpposite.unop_injective diff --git a/Mathlib/RingTheory/Ideal/AssociatedPrime.lean b/Mathlib/RingTheory/Ideal/AssociatedPrime.lean index e902fa6239cfa..0a316e3c547c5 100644 --- a/Mathlib/RingTheory/Ideal/AssociatedPrime.lean +++ b/Mathlib/RingTheory/Ideal/AssociatedPrime.lean @@ -161,7 +161,7 @@ theorem associatedPrimes.eq_singleton_of_isPrimary [IsNoetherianRing R] (hI : I. refine ⟨IsAssociatedPrime.eq_radical hI, ?_⟩ rintro rfl haveI : Nontrivial (R ⧸ I) := by - refine ⟨(Ideal.Quotient.mk I : _) 1, (Ideal.Quotient.mk I : _) 0, ?_⟩ + refine ⟨(Ideal.Quotient.mk I :) 1, (Ideal.Quotient.mk I :) 0, ?_⟩ rw [Ne, Ideal.Quotient.eq, sub_zero, ← Ideal.eq_top_iff_one] exact hI.1 obtain ⟨a, ha⟩ := associatedPrimes.nonempty R (R ⧸ I) diff --git a/Mathlib/RingTheory/Ideal/Basic.lean b/Mathlib/RingTheory/Ideal/Basic.lean index 9bcac7b686908..b2409062e4537 100644 --- a/Mathlib/RingTheory/Ideal/Basic.lean +++ b/Mathlib/RingTheory/Ideal/Basic.lean @@ -28,9 +28,7 @@ Support right ideals, and two-sided ideals over non-commutative rings. -/ -universe u v w - -variable {α : Type u} {β : Type v} {F : Type w} +variable {ι α β F : Type*} open Set Function @@ -40,20 +38,18 @@ section Semiring namespace Ideal -variable [Semiring α] (I : Ideal α) {a b : α} +variable {α : ι → Type*} [Π i, Semiring (α i)] (I : Π i, Ideal (α i)) section Pi -variable (ι : Type v) - -/-- `I^n` as an ideal of `R^n`. -/ -def pi : Ideal (ι → α) where - carrier := { x | ∀ i, x i ∈ I } - zero_mem' _i := I.zero_mem - add_mem' ha hb i := I.add_mem (ha i) (hb i) - smul_mem' a _b hb i := I.mul_mem_left (a i) (hb i) +/-- `Πᵢ Iᵢ` as an ideal of `Πᵢ Rᵢ`. -/ +def pi : Ideal (Π i, α i) where + carrier := { x | ∀ i, x i ∈ I i } + zero_mem' i := (I i).zero_mem + add_mem' ha hb i := (I i).add_mem (ha i) (hb i) + smul_mem' a _b hb i := (I i).mul_mem_left (a i) (hb i) -theorem mem_pi (x : ι → α) : x ∈ I.pi ι ↔ ∀ i, x i ∈ I := +theorem mem_pi (x : Π i, α i) : x ∈ pi I ↔ ∀ i, x i ∈ I i := Iff.rfl end Pi @@ -166,7 +162,7 @@ end CommSemiring section DivisionSemiring -variable {K : Type u} [DivisionSemiring K] (I : Ideal K) +variable {K : Type*} [DivisionSemiring K] (I : Ideal K) namespace Ideal @@ -254,7 +250,7 @@ end Ring namespace Ideal -variable {R : Type u} [CommSemiring R] [Nontrivial R] +variable {R : Type*} [CommSemiring R] [Nontrivial R] theorem bot_lt_of_maximal (M : Ideal R) [hm : M.IsMaximal] (non_field : ¬IsField R) : ⊥ < M := by rcases Ring.not_isField_iff_exists_ideal_bot_lt_and_lt_top.1 non_field with ⟨I, Ibot, Itop⟩ diff --git a/Mathlib/RingTheory/Ideal/Basis.lean b/Mathlib/RingTheory/Ideal/Basis.lean index ac26138204089..0639ab56e0dd4 100644 --- a/Mathlib/RingTheory/Ideal/Basis.lean +++ b/Mathlib/RingTheory/Ideal/Basis.lean @@ -39,8 +39,7 @@ theorem basisSpanSingleton_apply (b : Basis ι R S) {x : S} (hx : x ≠ 0) (i : theorem constr_basisSpanSingleton {N : Type*} [Semiring N] [Module N S] [SMulCommClass R N S] (b : Basis ι R S) {x : S} (hx : x ≠ 0) : (b.constr N).toFun (((↑) : _ → S) ∘ (basisSpanSingleton b hx)) = Algebra.lmul R S x := - b.ext fun i => by - erw [Basis.constr_basis, Function.comp_apply, basisSpanSingleton_apply, LinearMap.mul_apply'] + b.ext fun i => by simp end Ideal diff --git a/Mathlib/RingTheory/Ideal/Maps.lean b/Mathlib/RingTheory/Ideal/Maps.lean index a0cca39024554..16fd9d345f686 100644 --- a/Mathlib/RingTheory/Ideal/Maps.lean +++ b/Mathlib/RingTheory/Ideal/Maps.lean @@ -334,6 +334,22 @@ theorem IsMaximal.comap_piEvalRingHom {ι : Type*} {R : ι → Type*} [∀ i, Se · simpa [eq_comm] using eq · simp [update_of_ne ne] +theorem comap_le_comap_iff_of_surjective (hf : Function.Surjective f) (I J : Ideal S) : + comap f I ≤ comap f J ↔ I ≤ J := + ⟨fun h => (map_comap_of_surjective f hf I).symm.le.trans (map_le_of_le_comap h), fun h => + le_comap_of_map_le ((map_comap_of_surjective f hf I).le.trans h)⟩ + +/-- The map on ideals induced by a surjective map preserves inclusion. -/ +def orderEmbeddingOfSurjective (hf : Function.Surjective f) : Ideal S ↪o Ideal R where + toFun := comap f + inj' _ _ eq := SetLike.ext' (Set.preimage_injective.mpr hf <| SetLike.ext'_iff.mp eq) + map_rel_iff' := comap_le_comap_iff_of_surjective _ hf .. + +theorem map_eq_top_or_isMaximal_of_surjective (hf : Function.Surjective f) {I : Ideal R} + (H : IsMaximal I) : map f I = ⊤ ∨ IsMaximal (map f I) := + or_iff_not_imp_left.2 fun ne_top ↦ ⟨⟨ne_top, fun _J hJ ↦ comap_injective_of_surjective f hf <| + H.1.2 _ (le_comap_map.trans_lt <| (orderEmbeddingOfSurjective f hf).strictMono hJ)⟩⟩ + end Surjective section Injective @@ -438,6 +454,11 @@ instance map_isMaximal_of_equiv {E : Type*} [EquivLike E R S] [RingEquivClass E {p : Ideal R} [hp : p.IsMaximal] : (map e p).IsMaximal := hp.map_bijective e (EquivLike.bijective e) +/-- The pullback of a maximal ideal under a ring isomorphism is a maximal ideal. -/ +instance comap_isMaximal_of_equiv {E : Type*} [EquivLike E R S] [RingEquivClass E R S] (e : E) + {p : Ideal S} [hp : p.IsMaximal] : (comap e p).IsMaximal := + hp.comap_bijective e (EquivLike.bijective e) + theorem isMaximal_iff_of_bijective : (⊥ : Ideal R).IsMaximal ↔ (⊥ : Ideal S).IsMaximal := ⟨fun h ↦ map_bot (f := f) ▸ h.map_bijective f hf, fun h ↦ have e := RingEquiv.ofBijective f hf map_bot (f := e.symm) ▸ h.map_bijective _ e.symm.bijective⟩ @@ -481,19 +502,7 @@ def relIsoOfSurjective (hf : Function.Surjective f) : ⟨fun H => map_comap_of_surjective f hf I1 ▸ map_comap_of_surjective f hf I2 ▸ map_mono H, comap_mono⟩ -/-- The map on ideals induced by a surjective map preserves inclusion. -/ -def orderEmbeddingOfSurjective (hf : Function.Surjective f) : Ideal S ↪o Ideal R := - (relIsoOfSurjective f hf).toRelEmbedding.trans (Subtype.relEmbedding (fun x y => x ≤ y) _) - -theorem map_eq_top_or_isMaximal_of_surjective (hf : Function.Surjective f) {I : Ideal R} - (H : IsMaximal I) : map f I = ⊤ ∨ IsMaximal (map f I) := by - refine or_iff_not_imp_left.2 fun ne_top => ⟨⟨fun h => ne_top h, fun J hJ => ?_⟩⟩ - · refine - (relIsoOfSurjective f hf).injective - (Subtype.ext_iff.2 (Eq.trans (H.1.2 (comap f J) (lt_of_le_of_ne ?_ ?_)) comap_top.symm)) - · exact map_le_iff_le_comap.1 (le_of_lt hJ) - · exact fun h => hJ.right (le_map_of_comap_le_of_surjective f hf (le_of_eq h.symm)) - +-- Does not hold if `R` is a semiring: consider `ℕ → ZMod 2`. theorem comap_isMaximal_of_surjective (hf : Function.Surjective f) {K : Ideal S} [H : IsMaximal K] : IsMaximal (comap f K) := by refine ⟨⟨comap_ne_top _ H.1.1, fun J hJ => ?_⟩⟩ @@ -509,16 +518,6 @@ theorem comap_isMaximal_of_surjective (hf : Function.Surjective f) {K : Ideal S} rw [comap_map_of_surjective _ hf, sup_eq_left] exact le_trans (comap_mono bot_le) (le_of_lt hJ) -/-- The pullback of a maximal ideal under a ring isomorphism is a maximal ideal. -/ -instance comap_isMaximal_of_equiv {E : Type*} [EquivLike E R S] [RingEquivClass E R S] (e : E) - {p : Ideal S} [p.IsMaximal] : (comap e p).IsMaximal := - comap_isMaximal_of_surjective e (EquivLike.surjective e) - -theorem comap_le_comap_iff_of_surjective (hf : Function.Surjective f) (I J : Ideal S) : - comap f I ≤ comap f J ↔ I ≤ J := - ⟨fun h => (map_comap_of_surjective f hf I).symm.le.trans (map_le_of_le_comap h), fun h => - le_comap_of_map_le ((map_comap_of_surjective f hf I).le.trans h)⟩ - end Surjective @@ -669,7 +668,7 @@ theorem ker_rangeRestrict (f : R →+* S) : ker f.rangeRestrict = ker f := end RingRing /-- The kernel of a homomorphism to a domain is a prime ideal. -/ -theorem ker_isPrime {F : Type*} [Ring R] [Ring S] [IsDomain S] +theorem ker_isPrime {F : Type*} [Semiring R] [Semiring S] [IsDomain S] [FunLike F R S] [RingHomClass F R S] (f : F) : (ker f).IsPrime := ⟨by @@ -679,7 +678,7 @@ theorem ker_isPrime {F : Type*} [Ring R] [Ring S] [IsDomain S] simpa only [mem_ker, map_mul] using @eq_zero_or_eq_zero_of_mul_eq_zero S _ _ _ _ _⟩ /-- The kernel of a homomorphism to a field is a maximal ideal. -/ -theorem ker_isMaximal_of_surjective {R K F : Type*} [Ring R] [Field K] +theorem ker_isMaximal_of_surjective {R K F : Type*} [Ring R] [DivisionRing K] [FunLike F R K] [RingHomClass F R K] (f : F) (hf : Function.Surjective f) : (ker f).IsMaximal := by refine @@ -836,6 +835,10 @@ instance map_isPrime_of_equiv {F' : Type*} [EquivLike F' R S] [RingEquivClass F' rw [h, map_comap_of_equiv (f : R ≃+* S)] exact Ideal.IsPrime.comap (RingEquiv.symm (f : R ≃+* S)) +theorem map_eq_bot_iff_of_injective {I : Ideal R} {f : F} (hf : Function.Injective f) : + I.map f = ⊥ ↔ I = ⊥ := by + simp [map, span_eq_bot, ← map_zero f, -map_zero, hf.eq_iff, I.eq_bot_iff] + end Semiring section Ring @@ -882,10 +885,6 @@ theorem map_isPrime_of_surjective {f : F} (hf : Function.Surjective f) {I : Idea exact (H.mem_or_mem this).imp (fun h => ha ▸ mem_map_of_mem f h) fun h => hb ▸ mem_map_of_mem f h -theorem map_eq_bot_iff_of_injective {I : Ideal R} {f : F} (hf : Function.Injective f) : - I.map f = ⊥ ↔ I = ⊥ := by - rw [map_eq_bot_iff_le_ker, (RingHom.injective_iff_ker_eq_bot f).mp hf, le_bot_iff] - end Ring section CommRing @@ -1036,9 +1035,9 @@ theorem of_ker_algebraMap_eq_bot (R A : Type*) [CommRing R] [Semiring A] [Algebr [NoZeroDivisors A] (h : RingHom.ker (algebraMap R A) = ⊥) : NoZeroSMulDivisors R A := of_algebraMap_injective ((RingHom.injective_iff_ker_eq_bot _).mpr h) -theorem ker_algebraMap_eq_bot (R A : Type*) [CommRing R] [Ring A] [Nontrivial A] [Algebra R A] - [NoZeroSMulDivisors R A] : RingHom.ker (algebraMap R A) = ⊥ := - (RingHom.injective_iff_ker_eq_bot _).mp (algebraMap_injective R A) +theorem ker_algebraMap_eq_bot (R A : Type*) [CommSemiring R] [Semiring A] [Nontrivial A] + [Algebra R A] [NoZeroSMulDivisors R A] : RingHom.ker (algebraMap R A) = ⊥ := by + ext; simp [Algebra.algebraMap_eq_smul_one'] theorem iff_ker_algebraMap_eq_bot {R A : Type*} [CommRing R] [Ring A] [IsDomain A] [Algebra R A] : NoZeroSMulDivisors R A ↔ RingHom.ker (algebraMap R A) = ⊥ := diff --git a/Mathlib/RingTheory/Ideal/Operations.lean b/Mathlib/RingTheory/Ideal/Operations.lean index a04061874ab5e..3b8ded22c58e1 100644 --- a/Mathlib/RingTheory/Ideal/Operations.lean +++ b/Mathlib/RingTheory/Ideal/Operations.lean @@ -484,6 +484,7 @@ theorem multiset_prod_span_singleton (m : Multiset R) : Multiset.induction_on m (by simp) fun a m ih => by simp only [Multiset.map_cons, Multiset.prod_cons, ih, ← Ideal.span_singleton_mul_span_singleton] +open scoped Function in -- required for scoped `on` notation theorem finset_inf_span_singleton {ι : Type*} (s : Finset ι) (I : ι → R) (hI : Set.Pairwise (↑s) (IsCoprime on I)) : (s.inf fun i => Ideal.span ({I i} : Set R)) = Ideal.span {∏ i ∈ s, I i} := by @@ -851,6 +852,14 @@ variable (R) in theorem top_pow (n : ℕ) : (⊤ ^ n : Ideal R) = ⊤ := Nat.recOn n one_eq_top fun n ih => by rw [pow_succ, ih, top_mul] +theorem natCast_eq_top {n : ℕ} (hn : n ≠ 0) : (n : Ideal R) = ⊤ := + natCast_eq_one hn |>.trans one_eq_top + +/-- `3 : Ideal R` is *not* the ideal generated by 3 (which would be spelt + `Ideal.span {3}`), it is simply `1 + 1 + 1 = ⊤`. -/ +theorem ofNat_eq_top {n : ℕ} [n.AtLeastTwo] : (ofNat(n) : Ideal R) = ⊤ := + ofNat_eq_one.trans one_eq_top + variable (I) lemma radical_pow : ∀ {n}, n ≠ 0 → radical (I ^ n) = radical I diff --git a/Mathlib/RingTheory/Ideal/Prod.lean b/Mathlib/RingTheory/Ideal/Prod.lean index 0e2afa702d110..44e48d9060023 100644 --- a/Mathlib/RingTheory/Ideal/Prod.lean +++ b/Mathlib/RingTheory/Ideal/Prod.lean @@ -22,15 +22,7 @@ variable {R : Type u} {S : Type v} [Semiring R] [Semiring S] (I : Ideal R) (J : namespace Ideal /-- `I × J` as an ideal of `R × S`. -/ -def prod : Ideal (R × S) where - carrier := { x | x.fst ∈ I ∧ x.snd ∈ J } - zero_mem' := by simp - add_mem' := by - rintro ⟨a₁, a₂⟩ ⟨b₁, b₂⟩ ⟨ha₁, ha₂⟩ ⟨hb₁, hb₂⟩ - exact ⟨I.add_mem ha₁ hb₁, J.add_mem ha₂ hb₂⟩ - smul_mem' := by - rintro ⟨a₁, a₂⟩ ⟨b₁, b₂⟩ ⟨hb₁, hb₂⟩ - exact ⟨I.mul_mem_left _ hb₁, J.mul_mem_left _ hb₂⟩ +def prod : Ideal (R × S) := I.comap (RingHom.fst R S) ⊓ J.comap (RingHom.snd R S) @[simp] theorem mem_prod {r : R} {s : S} : (⟨r, s⟩ : R × S) ∈ prod I J ↔ r ∈ I ∧ s ∈ J := @@ -78,11 +70,16 @@ theorem map_prodComm_prod : /-- Ideals of `R × S` are in one-to-one correspondence with pairs of ideals of `R` and ideals of `S`. -/ -def idealProdEquiv : Ideal (R × S) ≃ Ideal R × Ideal S where +def idealProdEquiv : Ideal (R × S) ≃o Ideal R × Ideal S where toFun I := ⟨map (RingHom.fst R S) I, map (RingHom.snd R S) I⟩ invFun I := prod I.1 I.2 left_inv I := (ideal_prod_eq I).symm right_inv := fun ⟨I, J⟩ => by simp + map_rel_iff' {I J} := by + simp only [Equiv.coe_fn_mk, ge_iff_le, Prod.mk_le_mk] + refine ⟨fun h ↦ ?_, fun h ↦ ⟨map_mono h, map_mono h⟩⟩ + rw [ideal_prod_eq I, ideal_prod_eq J] + exact inf_le_inf (comap_mono h.1) (comap_mono h.2) @[simp] theorem idealProdEquiv_symm_apply (I : Ideal R) (J : Ideal S) : diff --git a/Mathlib/RingTheory/Ideal/Quotient/Basic.lean b/Mathlib/RingTheory/Ideal/Quotient/Basic.lean index 244da7d3aada9..dc04cdcc3ae36 100644 --- a/Mathlib/RingTheory/Ideal/Quotient/Basic.lean +++ b/Mathlib/RingTheory/Ideal/Quotient/Basic.lean @@ -151,55 +151,32 @@ section Pi variable (ι : Type v) /-- `R^n/I^n` is a `R/I`-module. -/ -instance modulePi : Module (R ⧸ I) ((ι → R) ⧸ I.pi ι) where +instance modulePi : Module (R ⧸ I) ((ι → R) ⧸ pi fun _ ↦ I) where smul c m := - Quotient.liftOn₂' c m (fun r m => Submodule.Quotient.mk <| r • m) <| by + Quotient.liftOn₂' c m (fun r m ↦ Submodule.Quotient.mk <| r • m) <| by intro c₁ m₁ c₂ m₂ hc hm apply Ideal.Quotient.eq.2 rw [Submodule.quotientRel_def] at hc hm intro i exact I.mul_sub_mul_mem hc (hm i) - one_smul := by - rintro ⟨a⟩ - convert_to Ideal.Quotient.mk (I.pi ι) _ = Ideal.Quotient.mk (I.pi ι) _ - congr with i; exact one_mul (a i) - mul_smul := by - rintro ⟨a⟩ ⟨b⟩ ⟨c⟩ - convert_to Ideal.Quotient.mk (I.pi ι) _ = Ideal.Quotient.mk (I.pi ι) _ - congr 1; funext i; exact mul_assoc a b (c i) - smul_add := by - rintro ⟨a⟩ ⟨b⟩ ⟨c⟩ - convert_to Ideal.Quotient.mk (I.pi ι) _ = Ideal.Quotient.mk (I.pi ι) _ - congr with i; exact mul_add a (b i) (c i) - smul_zero := by - rintro ⟨a⟩ - convert_to Ideal.Quotient.mk (I.pi ι) _ = Ideal.Quotient.mk (I.pi ι) _ - congr with _; exact mul_zero a - add_smul := by - rintro ⟨a⟩ ⟨b⟩ ⟨c⟩ - convert_to Ideal.Quotient.mk (I.pi ι) _ = Ideal.Quotient.mk (I.pi ι) _ - congr with i; exact add_mul a b (c i) - zero_smul := by - rintro ⟨a⟩ - convert_to Ideal.Quotient.mk (I.pi ι) _ = Ideal.Quotient.mk (I.pi ι) _ - congr with i; exact zero_mul (a i) + one_smul := by rintro ⟨a⟩; exact congr_arg _ (one_smul _ _) + mul_smul := by rintro ⟨a⟩ ⟨b⟩ ⟨c⟩; exact congr_arg _ (mul_smul _ _ _) + smul_add := by rintro ⟨a⟩ ⟨b⟩ ⟨c⟩; exact congr_arg _ (smul_add _ _ _) + smul_zero := by rintro ⟨a⟩; exact congr_arg _ (smul_zero _) + add_smul := by rintro ⟨a⟩ ⟨b⟩ ⟨c⟩; exact congr_arg _ (add_smul _ _ _) + zero_smul := by rintro ⟨a⟩; exact congr_arg _ (zero_smul _ _) /-- `R^n/I^n` is isomorphic to `(R/I)^n` as an `R/I`-module. -/ -noncomputable def piQuotEquiv : ((ι → R) ⧸ I.pi ι) ≃ₗ[R ⧸ I] ι → (R ⧸ I) where - toFun := fun x ↦ - Quotient.liftOn' x (fun f i => Ideal.Quotient.mk I (f i)) fun _ _ hab => - funext fun i => (Submodule.Quotient.eq' _).2 (QuotientAddGroup.leftRel_apply.mp hab i) +noncomputable def piQuotEquiv : ((ι → R) ⧸ pi fun _ ↦ I) ≃ₗ[R ⧸ I] ι → (R ⧸ I) where + toFun x := Quotient.liftOn' x (fun f i ↦ Ideal.Quotient.mk I (f i)) fun _ _ hab ↦ + funext fun i ↦ (Submodule.Quotient.eq' _).2 (QuotientAddGroup.leftRel_apply.mp hab i) map_add' := by rintro ⟨_⟩ ⟨_⟩; rfl map_smul' := by rintro ⟨_⟩ ⟨_⟩; rfl - invFun := fun x ↦ Ideal.Quotient.mk (I.pi ι) fun i ↦ Quotient.out (x i) + invFun x := Ideal.Quotient.mk _ (Quotient.out <| x ·) left_inv := by rintro ⟨x⟩ - exact Ideal.Quotient.eq.2 fun i => Ideal.Quotient.eq.1 (Quotient.out_eq' _) - right_inv := by - intro x - ext i - obtain ⟨_, _⟩ := @Quot.exists_rep _ _ (x i) - convert Quotient.out_eq' (x i) + exact Ideal.Quotient.eq.2 fun i ↦ Ideal.Quotient.eq.1 (Quotient.out_eq' _) + right_inv x := funext fun i ↦ Quotient.out_eq' (x i) /-- If `f : R^n → R^m` is an `R`-linear map and `I ⊆ R` is an ideal, then the image of `I^n` is contained in `I^m`. -/ diff --git a/Mathlib/RingTheory/Ideal/Quotient/Defs.lean b/Mathlib/RingTheory/Ideal/Quotient/Defs.lean index 5bd2896800b7f..a060399646f49 100644 --- a/Mathlib/RingTheory/Ideal/Quotient/Defs.lean +++ b/Mathlib/RingTheory/Ideal/Quotient/Defs.lean @@ -89,7 +89,7 @@ See note [partially-applied ext lemmas]. -/ @[ext 1100] theorem ringHom_ext [NonAssocSemiring S] ⦃f g : R ⧸ I →+* S⦄ (h : f.comp (mk I) = g.comp (mk I)) : f = g := - RingHom.ext fun x => Quotient.inductionOn' x <| (RingHom.congr_fun h : _) + RingHom.ext fun x => Quotient.inductionOn' x <| (RingHom.congr_fun h :) instance inhabited : Inhabited (R ⧸ I) := ⟨mk I 37⟩ diff --git a/Mathlib/RingTheory/Ideal/Quotient/Operations.lean b/Mathlib/RingTheory/Ideal/Quotient/Operations.lean index 488ffed5743b2..81c079b7988ec 100644 --- a/Mathlib/RingTheory/Ideal/Quotient/Operations.lean +++ b/Mathlib/RingTheory/Ideal/Quotient/Operations.lean @@ -409,7 +409,7 @@ theorem Quotient.liftₐ_apply (I : Ideal A) (f : A →ₐ[R₁] B) (hI : ∀ a theorem Quotient.liftₐ_comp (I : Ideal A) (f : A →ₐ[R₁] B) (hI : ∀ a : A, a ∈ I → f a = 0) : (Ideal.Quotient.liftₐ I f hI).comp (Ideal.Quotient.mkₐ R₁ I) = f := - AlgHom.ext fun _ => (Ideal.Quotient.lift_mk I (f : A →+* B) hI : _) + AlgHom.ext fun _ => (Ideal.Quotient.lift_mk I (f : A →+* B) hI :) theorem KerLift.map_smul (f : A →ₐ[R₁] B) (r : R₁) (x : A ⧸ (RingHom.ker f)) : f.kerLift (r • x) = r • f.kerLift x := by @@ -445,11 +445,6 @@ def quotientKerAlgEquivOfRightInverse {f : A →ₐ[R₁] B} {g : B → A} { RingHom.quotientKerEquivOfRightInverse hf, kerLiftAlg f with } -@[deprecated (since := "2024-02-27")] -alias quotientKerAlgEquivOfRightInverse.apply := quotientKerAlgEquivOfRightInverse_apply -@[deprecated (since := "2024-02-27")] -alias QuotientKerAlgEquivOfRightInverseSymm.apply := quotientKerAlgEquivOfRightInverse_symm_apply - /-- The **first isomorphism theorem** for algebras. -/ @[simps!] noncomputable def quotientKerAlgEquivOfSurjective {f : A →ₐ[R₁] B} (hf : Function.Surjective f) : diff --git a/Mathlib/RingTheory/IsTensorProduct.lean b/Mathlib/RingTheory/IsTensorProduct.lean index dccdc2035a370..addd18bf5d59f 100644 --- a/Mathlib/RingTheory/IsTensorProduct.lean +++ b/Mathlib/RingTheory/IsTensorProduct.lean @@ -504,7 +504,7 @@ theorem Algebra.pushoutDesc_left [Algebra.IsPushout R S R' S'] {A : Type*} [Semi theorem Algebra.lift_algHom_comp_left [Algebra.IsPushout R S R' S'] {A : Type*} [Semiring A] [Algebra R A] (f : S →ₐ[R] A) (g : R' →ₐ[R] A) (H) : (Algebra.pushoutDesc S' f g H).comp (toAlgHom R S S') = f := - AlgHom.ext fun x => (Algebra.pushoutDesc_left S' f g H x : _) + AlgHom.ext fun x => (Algebra.pushoutDesc_left S' f g H x :) @[simp] theorem Algebra.pushoutDesc_right [Algebra.IsPushout R S R' S'] {A : Type*} [Semiring A] @@ -519,7 +519,7 @@ theorem Algebra.pushoutDesc_right [Algebra.IsPushout R S R' S'] {A : Type*} [Sem theorem Algebra.lift_algHom_comp_right [Algebra.IsPushout R S R' S'] {A : Type*} [Semiring A] [Algebra R A] (f : S →ₐ[R] A) (g : R' →ₐ[R] A) (H) : (Algebra.pushoutDesc S' f g H).comp (toAlgHom R R' S') = g := - AlgHom.ext fun x => (Algebra.pushoutDesc_right S' f g H x : _) + AlgHom.ext fun x => (Algebra.pushoutDesc_right S' f g H x :) @[ext (iff := false)] theorem Algebra.IsPushout.algHom_ext [H : Algebra.IsPushout R S R' S'] {A : Type*} [Semiring A] @@ -532,7 +532,7 @@ theorem Algebra.IsPushout.algHom_ext [H : Algebra.IsPushout R S R' S'] {A : Type · intro s s' e rw [Algebra.smul_def, map_mul, map_mul, e] congr 1 - exact (AlgHom.congr_fun h₂ s : _) + exact (AlgHom.congr_fun h₂ s :) · intro s₁ s₂ e₁ e₂ rw [map_add, map_add, e₁, e₂] diff --git a/Mathlib/RingTheory/Kaehler/Basic.lean b/Mathlib/RingTheory/Kaehler/Basic.lean index d668bcecc861a..ca03981436858 100644 --- a/Mathlib/RingTheory/Kaehler/Basic.lean +++ b/Mathlib/RingTheory/Kaehler/Basic.lean @@ -207,8 +207,8 @@ def KaehlerDifferential.D : Derivation R S (Ω[S⁄R]) := -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 erw [← LinearMap.map_smul_of_tower (M₂ := Ω[S⁄R]), ← LinearMap.map_smul_of_tower (M₂ := Ω[S⁄R]), ← map_add, Ideal.toCotangent_eq, pow_two] - convert Submodule.mul_mem_mul (KaehlerDifferential.one_smul_sub_smul_one_mem_ideal R a : _) - (KaehlerDifferential.one_smul_sub_smul_one_mem_ideal R b : _) using 1 + convert Submodule.mul_mem_mul (KaehlerDifferential.one_smul_sub_smul_one_mem_ideal R a :) + (KaehlerDifferential.one_smul_sub_smul_one_mem_ideal R b :) using 1 simp only [AddSubgroupClass.coe_sub, Submodule.coe_add, Submodule.coe_mk, TensorProduct.tmul_mul_tmul, mul_sub, sub_mul, mul_comm b, Submodule.coe_smul_of_tower, smul_sub, TensorProduct.smul_tmul', smul_eq_mul, mul_one] @@ -599,7 +599,7 @@ noncomputable def KaehlerDifferential.quotKerTotalEquiv : obtain ⟨x, rfl⟩ := Submodule.mkQ_surjective _ x exact LinearMap.congr_fun - (KaehlerDifferential.derivationQuotKerTotal_lift_comp_linearCombination R S : _) x + (KaehlerDifferential.derivationQuotKerTotal_lift_comp_linearCombination R S :) x right_inv := by intro x obtain ⟨x, rfl⟩ := KaehlerDifferential.linearCombination_surjective R S x diff --git a/Mathlib/RingTheory/Kaehler/CotangentComplex.lean b/Mathlib/RingTheory/Kaehler/CotangentComplex.lean index 79c48f487358b..819a67cdc7c96 100644 --- a/Mathlib/RingTheory/Kaehler/CotangentComplex.lean +++ b/Mathlib/RingTheory/Kaehler/CotangentComplex.lean @@ -206,10 +206,7 @@ def Hom.sub (f g : Hom P P') : P.CotangentSpace →ₗ[S] P'.Cotangent := by map_one_eq_zero' := ?_ leibniz' := ?_ } · ext - simp only [LinearMap.coe_comp, LinearMap.coe_restrictScalars, Function.comp_apply, - Cotangent.val_mk, Cotangent.val_zero, Ideal.toCotangent_eq_zero] - erw [LinearMap.codRestrict_apply] - simp only [LinearMap.sub_apply, AlgHom.toLinearMap_apply, map_one, sub_self, Submodule.zero_mem] + simp [Ideal.toCotangent_eq_zero] · intro x y ext simp only [LinearMap.coe_comp, LinearMap.coe_restrictScalars, Function.comp_apply, @@ -377,7 +374,7 @@ lemma cotangentSpaceBasis_repr_one_tmul (x i) : simp lemma cotangentSpaceBasis_apply (i) : - P.cotangentSpaceBasis i = ((1 : S) ⊗ₜ[P.Ring] D R P.Ring (.X i) : _) := by + P.cotangentSpaceBasis i = ((1 : S) ⊗ₜ[P.Ring] D R P.Ring (.X i) :) := by simp [cotangentSpaceBasis, toExtension] universe w' u' v' diff --git a/Mathlib/RingTheory/Kaehler/JacobiZariski.lean b/Mathlib/RingTheory/Kaehler/JacobiZariski.lean index d4c5b8c6ef68f..6ccbe6af032c5 100644 --- a/Mathlib/RingTheory/Kaehler/JacobiZariski.lean +++ b/Mathlib/RingTheory/Kaehler/JacobiZariski.lean @@ -280,7 +280,7 @@ lemma δAux_C (r) : lemma δAux_toAlgHom {Q : Generators.{u₁} S T} {Q' : Generators.{u₃} S T} (f : Hom Q Q') (x) : δAux R Q' (f.toAlgHom x) = δAux R Q x + Finsupp.linearCombination _ (δAux R Q' ∘ f.val) - (Q.cotangentSpaceBasis.repr ((1 : T) ⊗ₜ[Q.Ring] D S Q.Ring x : _)) := by + (Q.cotangentSpaceBasis.repr ((1 : T) ⊗ₜ[Q.Ring] D S Q.Ring x :)) := by letI : AddCommGroup (T ⊗[S] Ω[S⁄R]) := inferInstance have : IsScalarTower Q.Ring Q.Ring T := IsScalarTower.left _ induction' x using MvPolynomial.induction_on with s x₁ x₂ hx₁ hx₂ p n IH diff --git a/Mathlib/RingTheory/Lasker.lean b/Mathlib/RingTheory/Lasker.lean index fc3785bfee908..77388bde77b56 100644 --- a/Mathlib/RingTheory/Lasker.lean +++ b/Mathlib/RingTheory/Lasker.lean @@ -55,6 +55,8 @@ lemma decomposition_erase_inf [DecidableEq (Ideal R)] {I : Ideal R} rw [← Finset.insert_erase hJ] at hs simp [← hs, hJ'] +open scoped Function -- required for scoped `on` notation + lemma isPrimary_decomposition_pairwise_ne_radical {I : Ideal R} {s : Finset (Ideal R)} (hs : s.inf id = I) (hs' : ∀ ⦃J⦄, J ∈ s → J.IsPrimary) : ∃ t : Finset (Ideal R), t.inf id = I ∧ (∀ ⦃J⦄, J ∈ t → J.IsPrimary) ∧ diff --git a/Mathlib/RingTheory/LocalProperties/Projective.lean b/Mathlib/RingTheory/LocalProperties/Projective.lean index bf3d45d29deed..dc9a8726d2715 100644 --- a/Mathlib/RingTheory/LocalProperties/Projective.lean +++ b/Mathlib/RingTheory/LocalProperties/Projective.lean @@ -121,7 +121,7 @@ theorem Module.projective_of_localization_maximal (H : ∀ (I : Ideal R) (_ : I. Module.Projective (Localization.AtPrime I) (LocalizedModule I.primeCompl M)) [Module.FinitePresentation R M] : Module.Projective R M := by have : Module.Finite R M := by infer_instance - have : (⊤ : Submodule R M).FG := this.out + have : (⊤ : Submodule R M).FG := this.fg_top have : ∃ (s : Finset M), _ := this obtain ⟨s, hs⟩ := this let N := s →₀ R diff --git a/Mathlib/RingTheory/LocalRing/Module.lean b/Mathlib/RingTheory/LocalRing/Module.lean index 8aec28a9bca3d..fdc9bee0bd9c1 100644 --- a/Mathlib/RingTheory/LocalRing/Module.lean +++ b/Mathlib/RingTheory/LocalRing/Module.lean @@ -61,7 +61,7 @@ theorem map_mkQ_eq {N₁ N₂ : Submodule R M} (h : N₁ ≤ N₂) (h' : N₂.FG theorem map_mkQ_eq_top {N : Submodule R M} [Module.Finite R M] : N.map (Submodule.mkQ (𝔪 • ⊤)) = ⊤ ↔ N = ⊤ := by - rw [← map_mkQ_eq (N₁ := N) le_top Module.Finite.out, Submodule.map_top, Submodule.range_mkQ] + rw [← map_mkQ_eq (N₁ := N) le_top Module.Finite.fg_top, Submodule.map_top, Submodule.range_mkQ] theorem map_tensorProduct_mk_eq_top {N : Submodule R M} [Module.Finite R M] : N.map (TensorProduct.mk R k M 1) = ⊤ ↔ N = ⊤ := by @@ -267,7 +267,7 @@ theorem free_of_lTensor_residueField_injective (hg : Surjective g) (h : Exact f (hf : Function.Injective (f.lTensor k)) : Module.Free R P := by have := Module.finitePresentation_of_free_of_surjective g hg - (by rw [h.linearMap_ker_eq, LinearMap.range_eq_map]; exact (Module.Finite.out).map f) + (by rw [h.linearMap_ker_eq, LinearMap.range_eq_map]; exact (Module.Finite.fg_top).map f) apply free_of_maximalIdeal_rTensor_injective rw [← LinearMap.lTensor_inj_iff_rTensor_inj] apply lTensor_injective_of_exact_of_exact_of_rTensor_injective diff --git a/Mathlib/RingTheory/Localization/Away/Basic.lean b/Mathlib/RingTheory/Localization/Away/Basic.lean index 97d2e44e131c1..d6ca20ff422c7 100644 --- a/Mathlib/RingTheory/Localization/Away/Basic.lean +++ b/Mathlib/RingTheory/Localization/Away/Basic.lean @@ -195,7 +195,6 @@ noncomputable def mapₐ (f : A →ₐ[R] B) (a : A) [Away a Aₚ] [Away (f a) B ⟨map Aₚ Bₚ f.toRingHom a, fun r ↦ by dsimp only [AlgHom.toRingHom_eq_coe, map, RingHom.coe_coe, OneHom.toFun_eq_coe] rw [IsScalarTower.algebraMap_apply R A Aₚ, IsScalarTower.algebraMap_eq R B Bₚ] - erw [IsLocalization.map_eq] simp⟩ @[simp] diff --git a/Mathlib/RingTheory/MvPolynomial/Homogeneous.lean b/Mathlib/RingTheory/MvPolynomial/Homogeneous.lean index 0ad41dd02719a..f0bcd26091b84 100644 --- a/Mathlib/RingTheory/MvPolynomial/Homogeneous.lean +++ b/Mathlib/RingTheory/MvPolynomial/Homogeneous.lean @@ -183,9 +183,6 @@ lemma _root_.MvPolynomial.isHomogeneous_C_mul_X (r : R) (i : σ) : (C r * X i).IsHomogeneous 1 := (isHomogeneous_X _ _).C_mul _ -@[deprecated (since := "2024-03-21")] -alias _root_.MvPolynomial.C_mul_X := _root_.MvPolynomial.isHomogeneous_C_mul_X - lemma pow (hφ : φ.IsHomogeneous m) (n : ℕ) : (φ ^ n).IsHomogeneous (m * n) := by rw [show φ ^ n = ∏ _i ∈ Finset.range n, φ by simp] rw [show m * n = ∑ _i ∈ Finset.range n, m by simp [mul_comm]] diff --git a/Mathlib/RingTheory/MvPowerSeries/Basic.lean b/Mathlib/RingTheory/MvPowerSeries/Basic.lean index 136414a51996a..1a27f73a40bef 100644 --- a/Mathlib/RingTheory/MvPowerSeries/Basic.lean +++ b/Mathlib/RingTheory/MvPowerSeries/Basic.lean @@ -347,7 +347,7 @@ def C : R →+* MvPowerSeries σ R := { monomial R (0 : σ →₀ ℕ) with map_one' := rfl map_mul' := fun a b => (monomial_mul_monomial 0 0 a b).symm - map_zero' := (monomial R (0 : _)).map_zero } + map_zero' := (monomial R 0).map_zero } variable {σ} {R} diff --git a/Mathlib/RingTheory/MvPowerSeries/Inverse.lean b/Mathlib/RingTheory/MvPowerSeries/Inverse.lean index be1a66314f9d2..87dca83d6c5c6 100644 --- a/Mathlib/RingTheory/MvPowerSeries/Inverse.lean +++ b/Mathlib/RingTheory/MvPowerSeries/Inverse.lean @@ -120,7 +120,7 @@ theorem mul_invOfUnit (φ : MvPowerSeries σ R) (u : Rˣ) (h : constantCoeff σ cases' hij with h₁ h₂ subst n rw [if_pos] - suffices (0 : _) + j < i + j by simpa + suffices 0 + j < i + j by simpa apply add_lt_add_right constructor · intro s diff --git a/Mathlib/RingTheory/Nakayama.lean b/Mathlib/RingTheory/Nakayama.lean index 0227b2a554422..ad4d01141b8f1 100644 --- a/Mathlib/RingTheory/Nakayama.lean +++ b/Mathlib/RingTheory/Nakayama.lean @@ -84,7 +84,7 @@ lemma eq_bot_of_set_smul_eq_of_subset_jacobson_annihilator {s : Set R} lemma top_ne_ideal_smul_of_le_jacobson_annihilator [Nontrivial M] [Module.Finite R M] {I} (h : I ≤ (Module.annihilator R M).jacobson) : (⊤ : Submodule R M) ≠ I • ⊤ := fun H => top_ne_bot <| - eq_bot_of_eq_ideal_smul_of_le_jacobson_annihilator Module.Finite.out H <| + eq_bot_of_eq_ideal_smul_of_le_jacobson_annihilator Module.Finite.fg_top H <| (congrArg (I ≤ Ideal.jacobson ·) annihilator_top).mpr h open Pointwise in diff --git a/Mathlib/RingTheory/Nilpotent/Lemmas.lean b/Mathlib/RingTheory/Nilpotent/Lemmas.lean index cfafb9cea8b59..b039f2c0a844d 100644 --- a/Mathlib/RingTheory/Nilpotent/Lemmas.lean +++ b/Mathlib/RingTheory/Nilpotent/Lemmas.lean @@ -103,7 +103,7 @@ lemma isNilpotent_restrict_of_le {f : End R M} {p q : Submodule R M} simp_rw [LinearMap.zero_apply, ZeroMemClass.coe_zero, ZeroMemClass.coe_eq_zero] at hn ⊢ rw [LinearMap.pow_restrict, LinearMap.restrict_apply] at hn ⊢ ext - exact (congr_arg Subtype.val hn : _) + exact (congr_arg Subtype.val hn :) lemma isNilpotent.restrict {f : M →ₗ[R] M} {p : Submodule R M} (hf : MapsTo f p p) (hnil : IsNilpotent f) : diff --git a/Mathlib/RingTheory/Noetherian/Basic.lean b/Mathlib/RingTheory/Noetherian/Basic.lean index 38156a143989d..6ddbaf53f56f5 100644 --- a/Mathlib/RingTheory/Noetherian/Basic.lean +++ b/Mathlib/RingTheory/Noetherian/Basic.lean @@ -7,7 +7,7 @@ import Mathlib.LinearAlgebra.Quotient.Basic import Mathlib.RingTheory.Noetherian.Defs import Mathlib.RingTheory.Finiteness.Cardinality import Mathlib.RingTheory.Finiteness.Finsupp -import Mathlib.RingTheory.Ideal.Maps +import Mathlib.RingTheory.Ideal.Prod /-! # Noetherian rings and modules @@ -56,7 +56,7 @@ open Set Pointwise section -variable {R : Type*} {M : Type*} {P : Type*} +variable {R M P : Type*} variable [Semiring R] [AddCommMonoid M] [AddCommMonoid P] variable [Module R M] [Module R P] @@ -133,9 +133,9 @@ end Module section -variable {R : Type*} {M : Type*} {P : Type*} -variable [Ring R] [AddCommGroup M] [AddCommGroup P] -variable [Module R M] [Module R P] +variable {R M N P : Type*} +variable [Ring R] [AddCommGroup M] [AddCommGroup N] [AddCommGroup P] +variable [Module R M] [Module R N] [Module R P] open IsNoetherian @@ -148,6 +148,7 @@ theorem fg_of_ker_bot [IsNoetherian R P] {N : Submodule R M} (f : M →ₗ[R] P) haveI := isNoetherian_of_ker_bot f hf IsNoetherian.noetherian N +-- False over a semiring: ℕ is a Noetherian ℕ-module but ℕ × ℕ is not. instance isNoetherian_prod [IsNoetherian R M] [IsNoetherian R P] : IsNoetherian R (M × P) := ⟨fun s => Submodule.fg_of_fg_map_of_fg_inf_ker (LinearMap.snd R M P) (noetherian _) <| @@ -164,7 +165,7 @@ variable {ι : Type*} [Finite ι] instance isNoetherian_pi : ∀ {M : ι → Type*} [∀ i, AddCommGroup (M i)] - [∀ i, Module R (M i)] [∀ i, IsNoetherian R (M i)], IsNoetherian R (∀ i, M i) := by + [∀ i, Module R (M i)] [∀ i, IsNoetherian R (M i)], IsNoetherian R (Π i, M i) := by apply Finite.induction_empty_option _ _ _ ι · exact fun e h ↦ isNoetherian_of_linearEquiv (LinearEquiv.piCongrLeft R _ e) · infer_instance @@ -183,6 +184,28 @@ instance isNoetherian_iSup : · intros; rw [iSup_of_empty]; infer_instance · intro _ _ ih _ _; rw [iSup_option]; infer_instance +/-- If the first and final modules in an exact sequence are Noetherian, + then the middle module is also Noetherian. -/ +theorem isNoetherian_of_range_eq_ker [IsNoetherian R M] [IsNoetherian R P] + (f : M →ₗ[R] N) (g : N →ₗ[R] P) (h : LinearMap.range f = LinearMap.ker g) : + IsNoetherian R N := + isNoetherian_mk <| + wellFounded_gt_exact_sequence + (LinearMap.range f) + (Submodule.map (f.ker.liftQ f le_rfl)) + (Submodule.comap (f.ker.liftQ f le_rfl)) + (Submodule.comap g.rangeRestrict) (Submodule.map g.rangeRestrict) + (Submodule.gciMapComap <| LinearMap.ker_eq_bot.mp <| Submodule.ker_liftQ_eq_bot _ _ _ le_rfl) + (Submodule.giMapComap g.surjective_rangeRestrict) + (by simp [Submodule.map_comap_eq, inf_comm, Submodule.range_liftQ]) + (by simp [Submodule.comap_map_eq, h]) + +theorem isNoetherian_iff_submodule_quotient (S : Submodule R P) : + IsNoetherian R P ↔ IsNoetherian R S ∧ IsNoetherian R (P ⧸ S) := by + refine ⟨fun _ ↦ ⟨inferInstance, inferInstance⟩, fun ⟨_, _⟩ ↦ ?_⟩ + apply isNoetherian_of_range_eq_ker S.subtype S.mkQ + rw [Submodule.ker_mkQ, Submodule.range_subtype] + end section CommRing @@ -204,10 +227,7 @@ open IsNoetherian Submodule Function section -universe w - -variable {R M P : Type*} {N : Type w} [Semiring R] [AddCommMonoid M] [Module R M] [AddCommMonoid N] - [Module R N] [AddCommMonoid P] [Module R P] +variable {R M : Type*} [Semiring R] [AddCommMonoid M] [Module R M] /-- If `∀ I > J, P I` implies `P J`, then `P` holds for all submodules. -/ theorem IsNoetherian.induction [IsNoetherian R M] {P : Submodule R M → Prop} @@ -218,10 +238,7 @@ end section -universe w - -variable {R M P : Type*} {N : Type w} [Ring R] [AddCommGroup M] [Module R M] [AddCommGroup N] - [Module R N] [AddCommGroup P] [Module R P] [IsNoetherian R M] +variable {R M N P : Type*} [Semiring R] [AddCommMonoid M] [Module R M] [IsNoetherian R M] lemma Submodule.finite_ne_bot_of_iSupIndep {ι : Type*} {N : ι → Submodule R M} (h : iSupIndep N) : @@ -246,28 +263,6 @@ theorem LinearIndependent.set_finite_of_isNoetherian [Nontrivial R] {s : Set M} (hi : LinearIndependent R ((↑) : s → M)) : s.Finite := @Set.toFinite _ _ hi.finite_of_isNoetherian -/-- If the first and final modules in an exact sequence are Noetherian, - then the middle module is also Noetherian. -/ -theorem isNoetherian_of_range_eq_ker [IsNoetherian R P] - (f : M →ₗ[R] N) (g : N →ₗ[R] P) (h : LinearMap.range f = LinearMap.ker g) : - IsNoetherian R N := - isNoetherian_mk <| - wellFounded_gt_exact_sequence - (LinearMap.range f) - (Submodule.map (f.ker.liftQ f le_rfl)) - (Submodule.comap (f.ker.liftQ f le_rfl)) - (Submodule.comap g.rangeRestrict) (Submodule.map g.rangeRestrict) - (Submodule.gciMapComap <| LinearMap.ker_eq_bot.mp <| Submodule.ker_liftQ_eq_bot _ _ _ le_rfl) - (Submodule.giMapComap g.surjective_rangeRestrict) - (by simp [Submodule.map_comap_eq, inf_comm, Submodule.range_liftQ]) - (by simp [Submodule.comap_map_eq, h]) - -theorem isNoetherian_iff_submodule_quotient (S : Submodule R P) : - IsNoetherian R P ↔ IsNoetherian R S ∧ IsNoetherian R (P ⧸ S) := by - refine ⟨fun _ ↦ ⟨inferInstance, inferInstance⟩, fun ⟨_, _⟩ ↦ ?_⟩ - apply isNoetherian_of_range_eq_ker S.subtype S.mkQ - rw [Submodule.ker_mkQ, Submodule.range_subtype] - /-- A sequence `f` of submodules of a noetherian module, with `f (n+1)` disjoint from the supremum of `f 0`, ..., `f n`, is eventually zero. -/ @@ -353,14 +348,30 @@ theorem isNoetherian_span_of_finite (R) {M} [Ring R] [AddCommGroup M] [Module R [IsNoetherianRing R] {A : Set M} (hA : A.Finite) : IsNoetherian R (Submodule.span R A) := isNoetherian_of_fg_of_noetherian _ (Submodule.fg_def.mpr ⟨A, hA, rfl⟩) -theorem isNoetherianRing_of_surjective (R) [Ring R] (S) [Ring S] (f : R →+* S) +theorem isNoetherianRing_of_surjective (R) [Semiring R] (S) [Semiring S] (f : R →+* S) (hf : Function.Surjective f) [H : IsNoetherianRing R] : IsNoetherianRing S := isNoetherian_mk ⟨OrderEmbedding.wellFounded (Ideal.orderEmbeddingOfSurjective f hf).dual H.wf⟩ -instance isNoetherianRing_range {R} [Ring R] {S} [Ring S] (f : R →+* S) [IsNoetherianRing R] : - IsNoetherianRing f.range := - isNoetherianRing_of_surjective R f.range f.rangeRestrict f.rangeRestrict_surjective +instance isNoetherianRing_rangeS {R} [Semiring R] {S} [Semiring S] (f : R →+* S) + [IsNoetherianRing R] : IsNoetherianRing f.rangeS := + isNoetherianRing_of_surjective R f.rangeS f.rangeSRestrict f.rangeSRestrict_surjective + +instance isNoetherianRing_range {R} [Ring R] {S} [Ring S] (f : R →+* S) + [IsNoetherianRing R] : IsNoetherianRing f.range := + isNoetherianRing_rangeS f -theorem isNoetherianRing_of_ringEquiv (R) [Ring R] {S} [Ring S] (f : R ≃+* S) [IsNoetherianRing R] : - IsNoetherianRing S := +theorem isNoetherianRing_of_ringEquiv (R) [Semiring R] {S} [Semiring S] (f : R ≃+* S) + [IsNoetherianRing R] : IsNoetherianRing S := isNoetherianRing_of_surjective R S f.toRingHom f.toEquiv.surjective + +instance {R S} [Semiring R] [Semiring S] [IsNoetherianRing R] [IsNoetherianRing S] : + IsNoetherianRing (R × S) := by + rw [IsNoetherianRing, isNoetherian_iff'] at * + exact Ideal.idealProdEquiv.toOrderEmbedding.wellFoundedGT + +instance {ι} [Finite ι] : ∀ {R : ι → Type*} [Π i, Semiring (R i)] [∀ i, IsNoetherianRing (R i)], + IsNoetherianRing (Π i, R i) := by + apply Finite.induction_empty_option _ _ _ ι + · exact fun e h ↦ isNoetherianRing_of_ringEquiv _ (.piCongrLeft _ e) + · infer_instance + · exact fun ih ↦ isNoetherianRing_of_ringEquiv _ (.symm .piOptionEquivProd) diff --git a/Mathlib/RingTheory/Noetherian/Nilpotent.lean b/Mathlib/RingTheory/Noetherian/Nilpotent.lean index da206c8a9170d..251104d3dbd8b 100644 --- a/Mathlib/RingTheory/Noetherian/Nilpotent.lean +++ b/Mathlib/RingTheory/Noetherian/Nilpotent.lean @@ -18,7 +18,7 @@ import Mathlib.RingTheory.Noetherian.Defs open IsNoetherian -theorem IsNoetherianRing.isNilpotent_nilradical (R : Type*) [CommRing R] [IsNoetherianRing R] : +theorem IsNoetherianRing.isNilpotent_nilradical (R : Type*) [CommSemiring R] [IsNoetherianRing R] : IsNilpotent (nilradical R) := by obtain ⟨n, hn⟩ := Ideal.exists_radical_pow_le_of_fg (⊥ : Ideal R) (IsNoetherian.noetherian _) exact ⟨n, eq_bot_iff.mpr hn⟩ diff --git a/Mathlib/RingTheory/NonUnitalSubring/Defs.lean b/Mathlib/RingTheory/NonUnitalSubring/Defs.lean index 557e930d7ef99..b0d51fc219306 100644 --- a/Mathlib/RingTheory/NonUnitalSubring/Defs.lean +++ b/Mathlib/RingTheory/NonUnitalSubring/Defs.lean @@ -159,7 +159,7 @@ theorem copy_eq (S : NonUnitalSubring R) (s : Set R) (hs : s = ↑S) : S.copy s theorem toNonUnitalSubsemiring_injective : Function.Injective (toNonUnitalSubsemiring : NonUnitalSubring R → NonUnitalSubsemiring R) - | _r, _s, h => ext (SetLike.ext_iff.mp h : _) + | _r, _s, h => ext (SetLike.ext_iff.mp h :) @[mono] theorem toNonUnitalSubsemiring_strictMono : @@ -173,7 +173,7 @@ theorem toNonUnitalSubsemiring_mono : theorem toAddSubgroup_injective : Function.Injective (toAddSubgroup : NonUnitalSubring R → AddSubgroup R) - | _r, _s, h => ext (SetLike.ext_iff.mp h : _) + | _r, _s, h => ext (SetLike.ext_iff.mp h :) @[mono] theorem toAddSubgroup_strictMono : @@ -185,7 +185,7 @@ theorem toAddSubgroup_mono : Monotone (toAddSubgroup : NonUnitalSubring R → Ad theorem toSubsemigroup_injective : Function.Injective (toSubsemigroup : NonUnitalSubring R → Subsemigroup R) - | _r, _s, h => ext (SetLike.ext_iff.mp h : _) + | _r, _s, h => ext (SetLike.ext_iff.mp h :) @[mono] theorem toSubsemigroup_strictMono : diff --git a/Mathlib/RingTheory/NonUnitalSubsemiring/Defs.lean b/Mathlib/RingTheory/NonUnitalSubsemiring/Defs.lean index 0dd80dd6b65cc..366fb8a79c8b3 100644 --- a/Mathlib/RingTheory/NonUnitalSubsemiring/Defs.lean +++ b/Mathlib/RingTheory/NonUnitalSubsemiring/Defs.lean @@ -119,11 +119,11 @@ theorem copy_eq (S : NonUnitalSubsemiring R) (s : Set R) (hs : s = ↑S) : S.cop theorem toSubsemigroup_injective : Function.Injective (toSubsemigroup : NonUnitalSubsemiring R → Subsemigroup R) - | _, _, h => ext (SetLike.ext_iff.mp h : _) + | _, _, h => ext (SetLike.ext_iff.mp h :) theorem toAddSubmonoid_injective : Function.Injective (toAddSubmonoid : NonUnitalSubsemiring R → AddSubmonoid R) - | _, _, h => ext (SetLike.ext_iff.mp h : _) + | _, _, h => ext (SetLike.ext_iff.mp h :) /-- Construct a `NonUnitalSubsemiring R` from a set `s`, a subsemigroup `sg`, and an additive submonoid `sa` such that `x ∈ s ↔ x ∈ sg ↔ x ∈ sa`. -/ diff --git a/Mathlib/RingTheory/Perfection.lean b/Mathlib/RingTheory/Perfection.lean index 7d7e45a3a1929..6dc12681fc5bd 100644 --- a/Mathlib/RingTheory/Perfection.lean +++ b/Mathlib/RingTheory/Perfection.lean @@ -247,7 +247,7 @@ variable {p R P} /-- A perfection map induces an isomorphism to the perfection. -/ noncomputable def equiv {π : P →+* R} (m : PerfectionMap p π) : P ≃+* Ring.Perfection R p := RingEquiv.ofBijective (Perfection.lift p P R π) - ⟨fun _ _ hxy => m.injective fun n => (congr_arg (Perfection.coeff R p n) hxy : _), fun f => + ⟨fun _ _ hxy => m.injective fun n => (congr_arg (Perfection.coeff R p n) hxy :), fun f => let ⟨x, hx⟩ := m.surjective f.1 f.2 ⟨x, Perfection.ext <| hx⟩⟩ diff --git a/Mathlib/RingTheory/PowerBasis.lean b/Mathlib/RingTheory/PowerBasis.lean index 1e67541ed7447..fc58485528703 100644 --- a/Mathlib/RingTheory/PowerBasis.lean +++ b/Mathlib/RingTheory/PowerBasis.lean @@ -75,7 +75,6 @@ theorem coe_basis (pb : PowerBasis R S) : ⇑pb.basis = fun i : Fin pb.dim => pb /-- Cannot be an instance because `PowerBasis` cannot be a class. -/ theorem finite (pb : PowerBasis R S) : Module.Finite R S := .of_basis pb.basis -@[deprecated (since := "2024-03-05")] alias finiteDimensional := PowerBasis.finite theorem finrank [StrongRankCondition R] (pb : PowerBasis R S) : Module.finrank R S = pb.dim := by diff --git a/Mathlib/RingTheory/PrincipalIdealDomain.lean b/Mathlib/RingTheory/PrincipalIdealDomain.lean index 5343a17c2253a..77a6f3b5770e5 100644 --- a/Mathlib/RingTheory/PrincipalIdealDomain.lean +++ b/Mathlib/RingTheory/PrincipalIdealDomain.lean @@ -303,12 +303,6 @@ theorem isMaximal_of_irreducible [CommRing R] [IsPrincipalIdealRing R] {p : R} refine (of_irreducible_mul hp).resolve_right (mt (fun hb => ?_) (not_le_of_lt hI)) erw [Ideal.span_singleton_le_span_singleton, IsUnit.mul_right_dvd hb]⟩⟩ -@[deprecated (since := "2024-02-12")] -protected alias irreducible_iff_prime := irreducible_iff_prime - -@[deprecated (since := "2024-02-12")] -protected alias associates_irreducible_iff_prime := associates_irreducible_iff_prime - variable [CommRing R] [IsDomain R] [IsPrincipalIdealRing R] section diff --git a/Mathlib/RingTheory/Radical.lean b/Mathlib/RingTheory/Radical.lean index 8b60b67c3beb9..a213f7b39bfec 100644 --- a/Mathlib/RingTheory/Radical.lean +++ b/Mathlib/RingTheory/Radical.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jineon Baek, Seewoo Lee -/ import Mathlib.Algebra.EuclideanDomain.Basic +import Mathlib.Algebra.Order.Group.Finset import Mathlib.RingTheory.Coprime.Basic import Mathlib.RingTheory.UniqueFactorizationDomain.NormalizedFactors diff --git a/Mathlib/RingTheory/RingHom/Unramified.lean b/Mathlib/RingTheory/RingHom/Unramified.lean new file mode 100644 index 0000000000000..1016996901edb --- /dev/null +++ b/Mathlib/RingTheory/RingHom/Unramified.lean @@ -0,0 +1,111 @@ +/- +Copyright (c) 2024 Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Andrew Yang +-/ +import Mathlib.RingTheory.Unramified.Locus +import Mathlib.RingTheory.LocalProperties.Basic + +/-! + +# The meta properties of unramified ring homomorphisms. + +-/ + +namespace RingHom + +variable {R : Type*} {S : Type*} [CommRing R] [CommRing S] + +/-- +A ring homomorphism `R →+* A` is formally unramified if `Ω[A⁄R]` is trivial. +See `Algebra.FormallyUnramified`. +-/ +@[algebraize Algebra.FormallyUnramified] +def FormallyUnramified (f : R →+* S) : Prop := + letI := f.toAlgebra + Algebra.FormallyUnramified R S + +lemma formallyUnramified_algebraMap [Algebra R S] : + (algebraMap R S).FormallyUnramified ↔ Algebra.FormallyUnramified R S := by + delta FormallyUnramified + congr! + exact Algebra.algebra_ext _ _ fun _ ↦ rfl + +namespace FormallyUnramified + +lemma stableUnderComposition : + StableUnderComposition FormallyUnramified := by + intros R S T _ _ _ f g _ _ + algebraize [f, g, g.comp f] + exact .comp R S T + +lemma respectsIso : + RespectsIso FormallyUnramified := by + refine stableUnderComposition.respectsIso ?_ + intros R S _ _ e + letI := e.toRingHom.toAlgebra + exact Algebra.FormallyUnramified.of_surjective (Algebra.ofId R S) e.surjective + +lemma isStableUnderBaseChange : + IsStableUnderBaseChange FormallyUnramified := by + refine .mk _ respectsIso ?_ + intros R S T _ _ _ _ _ h + show (algebraMap _ _).FormallyUnramified + rw [formallyUnramified_algebraMap] at h ⊢ + infer_instance + +lemma holdsForLocalizationAway : + HoldsForLocalizationAway FormallyUnramified := by + intros R S _ _ _ r _ + rw [formallyUnramified_algebraMap] + exact .of_isLocalization (.powers r) + +lemma ofLocalizationPrime : + OfLocalizationPrime FormallyUnramified := by + intros R S _ _ f H + algebraize [f] + rw [FormallyUnramified, ← Algebra.unramifiedLocus_eq_univ_iff, Set.eq_univ_iff_forall] + intro x + let Rₓ := Localization.AtPrime (x.asIdeal.comap f) + let Sₓ := Localization.AtPrime x.asIdeal + have := Algebra.FormallyUnramified.of_isLocalization (Rₘ := Rₓ) (x.asIdeal.comap f).primeCompl + letI : Algebra Rₓ Sₓ := (Localization.localRingHom _ _ _ rfl).toAlgebra + have : IsScalarTower R Rₓ Sₓ := .of_algebraMap_eq + fun x ↦ (Localization.localRingHom_to_map _ _ _ rfl x).symm + have : Algebra.FormallyUnramified Rₓ Sₓ := H _ _ + exact Algebra.FormallyUnramified.comp R Rₓ Sₓ + +lemma ofLocalizationSpanTarget : + OfLocalizationSpanTarget FormallyUnramified := by + intros R S _ _ f s hs H + algebraize [f] + rw [FormallyUnramified, ← Algebra.unramifiedLocus_eq_univ_iff, Set.eq_univ_iff_forall] + intro x + obtain ⟨r, hr, hrx⟩ : ∃ r ∈ s, x ∈ PrimeSpectrum.basicOpen r := by + simpa using (PrimeSpectrum.iSup_basicOpen_eq_top_iff'.mpr hs).ge + (TopologicalSpace.Opens.mem_top x) + refine Algebra.basicOpen_subset_unramifiedLocus_iff.mpr ?_ hrx + convert H ⟨r, hr⟩ + dsimp + rw [← algebraMap_toAlgebra f, ← IsScalarTower.algebraMap_eq, + formallyUnramified_algebraMap] + +lemma propertyIsLocal : + PropertyIsLocal FormallyUnramified := by + constructor + · intro R S _ _ f r R' S' _ _ _ _ _ _ H + algebraize [f, (algebraMap S S').comp f, IsLocalization.Away.map R' S' f r] + have := Algebra.FormallyUnramified.of_isLocalization (Rₘ := S') (.powers (f r)) + have := Algebra.FormallyUnramified.comp R S S' + have H : Submonoid.powers r ≤ (Submonoid.powers (f r)).comap f := by + rintro x ⟨n, rfl⟩; exact ⟨n, by simp⟩ + have : IsScalarTower R R' S' := .of_algebraMap_eq' (IsLocalization.map_comp H).symm + exact Algebra.FormallyUnramified.of_comp R R' S' + · exact ofLocalizationSpanTarget + · exact ofLocalizationSpanTarget.ofLocalizationSpan + (stableUnderComposition.stableUnderCompositionWithLocalizationAway + holdsForLocalizationAway).1 + · exact (stableUnderComposition.stableUnderCompositionWithLocalizationAway + holdsForLocalizationAway).2 + +end RingHom.FormallyUnramified diff --git a/Mathlib/RingTheory/SimpleModule.lean b/Mathlib/RingTheory/SimpleModule.lean index dbae9bcb743ed..1561b18e289da 100644 --- a/Mathlib/RingTheory/SimpleModule.lean +++ b/Mathlib/RingTheory/SimpleModule.lean @@ -176,9 +176,6 @@ theorem IsSemisimpleModule.of_sSup_simples_eq_top (h : sSup { m : Submodule R M | IsSimpleModule R m } = ⊤) : IsSemisimpleModule R M := complementedLattice_of_sSup_atoms_eq_top (by simp_rw [← h, isSimpleModule_iff_isAtom]) -@[deprecated (since := "2024-03-05")] -alias is_semisimple_of_sSup_simples_eq_top := IsSemisimpleModule.of_sSup_simples_eq_top - namespace IsSemisimpleModule variable [IsSemisimpleModule R M] @@ -261,9 +258,6 @@ theorem sSup_simples_eq_top_iff_isSemisimpleModule : sSup { m : Submodule R M | IsSimpleModule R m } = ⊤ ↔ IsSemisimpleModule R M := ⟨.of_sSup_simples_eq_top, fun _ ↦ IsSemisimpleModule.sSup_simples_eq_top _ _⟩ -@[deprecated (since := "2024-03-05")] -alias is_semisimple_iff_top_eq_sSup_simples := sSup_simples_eq_top_iff_isSemisimpleModule - /-- A module generated by semisimple submodules is itself semisimple. -/ lemma isSemisimpleModule_of_isSemisimpleModule_submodule {s : Set ι} {p : ι → Submodule R M} (hp : ∀ i ∈ s, IsSemisimpleModule R (p i)) (hp' : ⨆ i ∈ s, p i = ⊤) : @@ -309,12 +303,12 @@ instance IsSemisimpleRing.isCoatomic_submodule [IsSemisimpleRing R] : IsCoatomic open LinearMap in /-- A finite product of semisimple rings is semisimple. -/ -instance {ι} [Finite ι] (R : ι → Type*) [∀ i, Ring (R i)] [∀ i, IsSemisimpleRing (R i)] : - IsSemisimpleRing (∀ i, R i) := by - letI (i) : Module (∀ i, R i) (R i) := Module.compHom _ (Pi.evalRingHom R i) +instance {ι} [Finite ι] (R : ι → Type*) [Π i, Ring (R i)] [∀ i, IsSemisimpleRing (R i)] : + IsSemisimpleRing (Π i, R i) := by + letI (i) : Module (Π i, R i) (R i) := Module.compHom _ (Pi.evalRingHom R i) let e (i) : R i →ₛₗ[Pi.evalRingHom R i] R i := { AddMonoidHom.id (R i) with map_smul' := fun _ _ ↦ rfl } - have (i) : IsSemisimpleModule (∀ i, R i) (R i) := + have (i) : IsSemisimpleModule (Π i, R i) (R i) := ((e i).isSemisimpleModule_iff_of_bijective Function.bijective_id).mpr inferInstance classical exact isSemisimpleModule_of_isSemisimpleModule_submodule' (p := (range <| single _ _ ·)) @@ -365,8 +359,8 @@ universe u in /-- The existence part of the Artin–Wedderburn theorem. -/ proof_wanted isSemisimpleRing_iff_pi_matrix_divisionRing {R : Type u} [Ring R] : IsSemisimpleRing R ↔ - ∃ (n : ℕ) (S : Fin n → Type u) (d : Fin n → ℕ) (_ : ∀ i, DivisionRing (S i)), - Nonempty (R ≃+* ∀ i, Matrix (Fin (d i)) (Fin (d i)) (S i)) + ∃ (n : ℕ) (S : Fin n → Type u) (d : Fin n → ℕ) (_ : Π i, DivisionRing (S i)), + Nonempty (R ≃+* Π i, Matrix (Fin (d i)) (Fin (d i)) (S i)) variable {ι R} diff --git a/Mathlib/RingTheory/Smooth/Basic.lean b/Mathlib/RingTheory/Smooth/Basic.lean index 4a40e5881dd41..0d42a4a3c9af6 100644 --- a/Mathlib/RingTheory/Smooth/Basic.lean +++ b/Mathlib/RingTheory/Smooth/Basic.lean @@ -98,7 +98,7 @@ theorem comp_lift [FormallySmooth R A] (I : Ideal B) (hI : IsNilpotent I) @[simp] theorem mk_lift [FormallySmooth R A] (I : Ideal B) (hI : IsNilpotent I) (g : A →ₐ[R] B ⧸ I) (x : A) : Ideal.Quotient.mk I (FormallySmooth.lift I hI g x) = g x := - AlgHom.congr_fun (FormallySmooth.comp_lift I hI g : _) x + AlgHom.congr_fun (FormallySmooth.comp_lift I hI g :) x variable {C : Type u} [CommRing C] [Algebra R C] @@ -202,7 +202,7 @@ theorem of_split [FormallySmooth R P] (g : A →ₐ[R] P ⧸ (RingHom.ker f.toRi refine Ideal.Quotient.liftₐ _ (FormallySmooth.lift I ⟨2, hI⟩ (i.comp f)) ?_ have : RingHom.ker f ≤ I.comap (FormallySmooth.lift I ⟨2, hI⟩ (i.comp f)) := by rintro x (hx : f x = 0) - have : _ = i (f x) := (FormallySmooth.mk_lift I ⟨2, hI⟩ (i.comp f) x : _) + have : _ = i (f x) := (FormallySmooth.mk_lift I ⟨2, hI⟩ (i.comp f) x :) rwa [hx, map_zero, ← Ideal.Quotient.mk_eq_mk, Submodule.Quotient.mk_eq_zero] at this intro x hx have := (Ideal.pow_right_mono this 2).trans (Ideal.le_comap_pow _ 2) hx diff --git a/Mathlib/RingTheory/Smooth/Kaehler.lean b/Mathlib/RingTheory/Smooth/Kaehler.lean index 780d71c547b67..47603d7179636 100644 --- a/Mathlib/RingTheory/Smooth/Kaehler.lean +++ b/Mathlib/RingTheory/Smooth/Kaehler.lean @@ -380,6 +380,8 @@ include hf in Given a formally smooth `R`-algebra `P` and a surjective algebra homomorphism `f : P →ₐ[R] S` with kernel `I` (typically a presentation `R[X] → S`), `S` is formally smooth iff the `P`-linear map `I/I² → S ⊗[P] Ω[P⁄R]` is split injective. +Also see `Algebra.Extension.formallySmooth_iff_split_injection` +for the version in terms of `Extension`. -/ @[stacks 031I] theorem Algebra.FormallySmooth.iff_split_injection : @@ -388,6 +390,26 @@ theorem Algebra.FormallySmooth.iff_split_injection : simp only [nonempty_subtype] at this rw [this, ← Algebra.FormallySmooth.iff_split_surjection _ hf] +/-- +Given a formally smooth `R`-algebra `P` and a surjective algebra homomorphism `f : P →ₐ[R] S` +with kernel `I` (typically a presentation `R[X] → S`), +`S` is formally smooth iff the `P`-linear map `I/I² → S ⊗[P] Ω[P⁄R]` is split injective. +-/ +@[stacks 031I] +theorem Algebra.Extension.formallySmooth_iff_split_injection + (P : Algebra.Extension.{u} R S) [FormallySmooth R P.Ring] : + Algebra.FormallySmooth R S ↔ ∃ l, l ∘ₗ P.cotangentComplex = LinearMap.id := by + refine (Algebra.FormallySmooth.iff_split_injection P.algebraMap_surjective).trans ?_ + let e : P.ker.Cotangent ≃ₗ[P.Ring] P.Cotangent := + { __ := AddEquiv.refl _, map_smul' r m := by ext1; simp; rfl } + constructor + · intro ⟨l, hl⟩ + exact ⟨(e.comp l).extendScalarsOfSurjective P.algebraMap_surjective, + LinearMap.ext (DFunLike.congr_fun hl : _)⟩ + · intro ⟨l, hl⟩ + exact ⟨e.symm.toLinearMap ∘ₗ l.restrictScalars P.Ring, + LinearMap.ext (DFunLike.congr_fun hl : _)⟩ + include hf in /-- Given a formally smooth `R`-algebra `P` and a surjective algebra homomorphism `f : P →ₐ[R] S` @@ -447,3 +469,6 @@ theorem Algebra.FormallySmooth.iff_subsingleton_and_projective : show Function.Injective (Generators.self R S).toExtension.cotangentComplex ↔ _ rw [← LinearMap.ker_eq_bot, ← Submodule.subsingleton_iff_eq_bot] rfl + +instance [Algebra.FormallySmooth R S] : Subsingleton (Algebra.H1Cotangent R S) := + (Algebra.FormallySmooth.iff_subsingleton_and_projective.mp ‹_›).1 diff --git a/Mathlib/RingTheory/Smooth/Local.lean b/Mathlib/RingTheory/Smooth/Local.lean new file mode 100644 index 0000000000000..589536e0c64de --- /dev/null +++ b/Mathlib/RingTheory/Smooth/Local.lean @@ -0,0 +1,37 @@ +/- +Copyright (c) 2024 Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Andrew Yang +-/ +import Mathlib.RingTheory.LocalRing.Module +import Mathlib.RingTheory.Smooth.Kaehler +import Mathlib.RingTheory.TensorProduct.Free + +/-! +# Formally smooth local algebras +-/ + +open TensorProduct IsLocalRing KaehlerDifferential + +/-- +The **Jacobian criterion** for smoothness of local algebras. +Suppose `S` is a local `R`-algebra, and `0 → I → P → S → 0` is a presentation such that +`P` is formally-smooth over `R`, `Ω[P⁄R]` is finite free over `P`, +(typically satisfied when `P` is the localization of a polynomial ring of finite type) +and `I` is finitely generated. +Then `S` is formally smooth iff `k ⊗ₛ I/I² → k ⊗ₚ Ω[P/R]` is injective, +where `k` is the residue field of `S`. +-/ +theorem Algebra.FormallySmooth.iff_injective_lTensor_residueField {R S} [CommRing R] + [CommRing S] [IsLocalRing S] [Algebra R S] (P : Algebra.Extension R S) + [FormallySmooth R P.Ring] + [Module.Free P.Ring (Ω[P.Ring⁄R])] [Module.Finite P.Ring (Ω[P.Ring⁄R])] + (h' : P.ker.FG) : + Algebra.FormallySmooth R S ↔ + Function.Injective (P.cotangentComplex.lTensor (ResidueField S)) := by + have : Module.Finite P.Ring P.Cotangent := + have : Module.Finite P.Ring P.ker := ⟨(Submodule.fg_top _).mpr h'⟩ + .of_surjective _ Extension.Cotangent.mk_surjective + have : Module.Finite S P.Cotangent := Module.Finite.of_restrictScalars_finite P.Ring _ _ + rw [← IsLocalRing.split_injective_iff_lTensor_residueField_injective, + P.formallySmooth_iff_split_injection] diff --git a/Mathlib/RingTheory/Smooth/Pi.lean b/Mathlib/RingTheory/Smooth/Pi.lean index cfb4cf1dd16df..aab45da1545e8 100644 --- a/Mathlib/RingTheory/Smooth/Pi.lean +++ b/Mathlib/RingTheory/Smooth/Pi.lean @@ -118,4 +118,7 @@ theorem pi_iff [Finite I] : rwa [mul_sub, ← map_mul, mul_comm, mul_assoc, (he.idem i).eq, he', ← map_mul, ← Pi.single_mul, one_mul, sub_eq_zero] at this +instance [Finite I] [∀ i, FormallySmooth R (A i)] : FormallySmooth R (Π i, A i) := + (pi_iff _).mpr ‹_› + end Algebra.FormallySmooth diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain/Defs.lean b/Mathlib/RingTheory/UniqueFactorizationDomain/Defs.lean index 9456594b84a19..5a7f8e147aaf6 100644 --- a/Mathlib/RingTheory/UniqueFactorizationDomain/Defs.lean +++ b/Mathlib/RingTheory/UniqueFactorizationDomain/Defs.lean @@ -135,11 +135,6 @@ instance (priority := 100) ufm_of_decomposition_of_wfDvdMonoid UniqueFactorizationMonoid α := { ‹WfDvdMonoid α› with irreducible_iff_prime := irreducible_iff_prime } -@[deprecated ufm_of_decomposition_of_wfDvdMonoid (since := "2024-02-12")] -theorem ufm_of_gcd_of_wfDvdMonoid [CancelCommMonoidWithZero α] [WfDvdMonoid α] - [DecompositionMonoid α] : UniqueFactorizationMonoid α := - ufm_of_decomposition_of_wfDvdMonoid - end Prio namespace UniqueFactorizationMonoid diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain/FactorSet.lean b/Mathlib/RingTheory/UniqueFactorizationDomain/FactorSet.lean index 2dd1d63b54515..d2a4e0e814f99 100644 --- a/Mathlib/RingTheory/UniqueFactorizationDomain/FactorSet.lean +++ b/Mathlib/RingTheory/UniqueFactorizationDomain/FactorSet.lean @@ -219,7 +219,6 @@ noncomputable def factors (a : Associates α) : FactorSet α := by theorem factors_zero : (0 : Associates α).factors = ⊤ := dif_pos rfl -@[deprecated (since := "2024-03-16")] alias factors_0 := factors_zero @[simp] theorem factors_mk (a : α) (h : a ≠ 0) : (Associates.mk a).factors = factors' a := by diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain/Multiplicity.lean b/Mathlib/RingTheory/UniqueFactorizationDomain/Multiplicity.lean index 9768ee91a10df..fb49eb4d8fc1f 100644 --- a/Mathlib/RingTheory/UniqueFactorizationDomain/Multiplicity.lean +++ b/Mathlib/RingTheory/UniqueFactorizationDomain/Multiplicity.lean @@ -128,9 +128,4 @@ theorem count_normalizedFactors_eq' [DecidableEq R] {p x : R} (hp : p = 0 ∨ Ir end multiplicity -/-- Deprecated. Use `WfDvdMonoid.max_power_factor` instead. -/ -@[deprecated WfDvdMonoid.max_power_factor (since := "2024-03-01")] -theorem max_power_factor {a₀ x : R} (h : a₀ ≠ 0) (hx : Irreducible x) : - ∃ n : ℕ, ∃ a : R, ¬x ∣ a ∧ a₀ = x ^ n * a := WfDvdMonoid.max_power_factor h hx - end UniqueFactorizationMonoid diff --git a/Mathlib/RingTheory/Unramified/Basic.lean b/Mathlib/RingTheory/Unramified/Basic.lean index 81416a42263dc..9e91002b6692d 100644 --- a/Mathlib/RingTheory/Unramified/Basic.lean +++ b/Mathlib/RingTheory/Unramified/Basic.lean @@ -39,8 +39,7 @@ namespace Algebra section -variable (R : Type v) [CommRing R] -variable (A : Type u) [CommRing A] [Algebra R A] +variable (R : Type v) (A : Type u) [CommRing R] [CommRing A] [Algebra R A] /-- An `R`-algebra `A` is formally unramified if `Ω[A⁄R]` is trivial. diff --git a/Mathlib/RingTheory/Unramified/Locus.lean b/Mathlib/RingTheory/Unramified/Locus.lean new file mode 100644 index 0000000000000..ec9cd9995a51e --- /dev/null +++ b/Mathlib/RingTheory/Unramified/Locus.lean @@ -0,0 +1,62 @@ +/- +Copyright (c) 2024 Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Andrew Yang +-/ +import Mathlib.AlgebraicGeometry.PrimeSpectrum.Basic +import Mathlib.RingTheory.Etale.Kaehler +import Mathlib.RingTheory.Support + +/-! +# Unramified locus of an algebra + +## Main results +- `Algebra.unramifiedLocus` : The set of primes that is unramified over the base. +- `Algebra.basicOpen_subset_unramifiedLocus_iff` : + `D(f)` is contained in the unramified locus if and only if `A_f` is unramified over `R`. +- `Algebra.unramifiedLocus_eq_univ_iff` : + The unramified locus is the whole spectrum if and only if `A` is unramified over `R`. +- `Algebra.isOpen_unramifiedLocus` : + If `A` is (essentially) of finite type over `R`, then the unramified locus is open. +-/ + +universe u + +variable (R A : Type u) [CommRing R] [CommRing A] [Algebra R A] + +namespace Algebra + +/-- `Algebra.unramifiedLocus R A` is the set of primes `p` of `A` +such that `Aₚ` is formally unramified over `R`. -/ +def unramifiedLocus : Set (PrimeSpectrum A) := + { p | Algebra.FormallyUnramified R (Localization.AtPrime p.asIdeal) } + +variable {R A} + +lemma unramifiedLocus_eq_compl_support : + unramifiedLocus R A = (Module.support A (Ω[A⁄R]))ᶜ := by + ext p + simp only [Set.mem_compl_iff, Module.not_mem_support_iff] + have := IsLocalizedModule.iso p.asIdeal.primeCompl + (KaehlerDifferential.map R R A (Localization.AtPrime p.asIdeal)) + exact (Algebra.formallyUnramified_iff _ _).trans this.subsingleton_congr.symm + +lemma basicOpen_subset_unramifiedLocus_iff {f : A} : + ↑(PrimeSpectrum.basicOpen f) ⊆ unramifiedLocus R A ↔ + Algebra.FormallyUnramified R (Localization.Away f) := by + rw [unramifiedLocus_eq_compl_support, Set.subset_compl_comm, + PrimeSpectrum.basicOpen_eq_zeroLocus_compl, compl_compl, + ← LocalizedModule.subsingleton_iff_support_subset, Algebra.formallyUnramified_iff] + exact (IsLocalizedModule.iso (.powers f) + (KaehlerDifferential.map R R A (Localization.Away f))).subsingleton_congr + +lemma unramifiedLocus_eq_univ_iff : + unramifiedLocus R A = Set.univ ↔ Algebra.FormallyUnramified R A := by + rw [unramifiedLocus_eq_compl_support, compl_eq_comm, Set.compl_univ, eq_comm, + Module.support_eq_empty_iff, Algebra.formallyUnramified_iff] + +lemma isOpen_unramifiedLocus [EssFiniteType R A] : IsOpen (unramifiedLocus R A) := by + rw [unramifiedLocus_eq_compl_support, Module.support_eq_zeroLocus] + exact (PrimeSpectrum.isClosed_zeroLocus _).isOpen_compl + +end Algebra diff --git a/Mathlib/RingTheory/Unramified/Pi.lean b/Mathlib/RingTheory/Unramified/Pi.lean index 755c16a601f58..ddebc82f3785d 100644 --- a/Mathlib/RingTheory/Unramified/Pi.lean +++ b/Mathlib/RingTheory/Unramified/Pi.lean @@ -95,4 +95,7 @@ theorem pi_iff : Ideal.Quotient.mkₐ_eq_mk] exact Ideal.mem_map_of_mem (Ideal.Quotient.mk J') (hf (Pi.single x r)) +instance [∀ i, FormallyUnramified R (f i)] : FormallyUnramified R (Π i, f i) := + (pi_iff _).mpr ‹_› + end Algebra.FormallyUnramified diff --git a/Mathlib/RingTheory/WittVector/Basic.lean b/Mathlib/RingTheory/WittVector/Basic.lean index e2c5bfed3057c..9258b5955fdb9 100644 --- a/Mathlib/RingTheory/WittVector/Basic.lean +++ b/Mathlib/RingTheory/WittVector/Basic.lean @@ -70,7 +70,7 @@ namespace mapFun theorem injective (f : α → β) (hf : Injective f) : Injective (mapFun f : 𝕎 α → 𝕎 β) := by intros _ _ h ext p - exact hf (congr_arg (fun x => coeff x p) h : _) + exact hf (congr_arg (fun x => coeff x p) h :) theorem surjective (f : α → β) (hf : Surjective f) : Surjective (mapFun f : 𝕎 α → 𝕎 β) := fun x => ⟨mk _ fun n => Classical.choose <| hf <| x.coeff n, diff --git a/Mathlib/SetTheory/Cardinal/Aleph.lean b/Mathlib/SetTheory/Cardinal/Aleph.lean index 896cc181244ef..9e69e6a4684de 100644 --- a/Mathlib/SetTheory/Cardinal/Aleph.lean +++ b/Mathlib/SetTheory/Cardinal/Aleph.lean @@ -139,9 +139,8 @@ theorem preOmega_natCast (n : ℕ) : preOmega n = n := by rw [Nat.cast_lt] exact lt_succ n --- See note [no_index around OfNat.ofNat] @[simp] -theorem preOmega_ofNat (n : ℕ) [n.AtLeastTwo] : preOmega (no_index (OfNat.ofNat n)) = n := +theorem preOmega_ofNat (n : ℕ) [n.AtLeastTwo] : preOmega ofNat(n) = n := preOmega_natCast n theorem preOmega_le_of_forall_lt {o a : Ordinal} (ha : IsInitial a) (H : ∀ b < o, preOmega b < a) : diff --git a/Mathlib/SetTheory/Cardinal/Arithmetic.lean b/Mathlib/SetTheory/Cardinal/Arithmetic.lean index 5a7416afca6f0..59acdc7b0b398 100644 --- a/Mathlib/SetTheory/Cardinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Cardinal/Arithmetic.lean @@ -463,16 +463,10 @@ theorem add_le_add_iff_of_lt_aleph0 {α β γ : Cardinal} (γ₀ : γ < ℵ₀) theorem add_nat_le_add_nat_iff {α β : Cardinal} (n : ℕ) : α + n ≤ β + n ↔ α ≤ β := add_le_add_iff_of_lt_aleph0 (nat_lt_aleph0 n) -@[deprecated (since := "2024-02-12")] -alias add_nat_le_add_nat_iff_of_lt_aleph_0 := add_nat_le_add_nat_iff - @[simp] theorem add_one_le_add_one_iff {α β : Cardinal} : α + 1 ≤ β + 1 ↔ α ≤ β := add_le_add_iff_of_lt_aleph0 one_lt_aleph0 -@[deprecated (since := "2024-02-12")] -alias add_one_le_add_one_iff_of_lt_aleph_0 := add_one_le_add_one_iff - end aleph /-! ### Properties about `power` -/ diff --git a/Mathlib/SetTheory/Cardinal/Basic.lean b/Mathlib/SetTheory/Cardinal/Basic.lean index e2bee0bf6130e..ca3df047a7bdc 100644 --- a/Mathlib/SetTheory/Cardinal/Basic.lean +++ b/Mathlib/SetTheory/Cardinal/Basic.lean @@ -1290,10 +1290,9 @@ theorem mk_fin (n : ℕ) : #(Fin n) = n := by simp @[simp] theorem lift_natCast (n : ℕ) : lift.{u} (n : Cardinal.{v}) = n := by induction n <;> simp [*] --- See note [no_index around OfNat.ofNat] @[simp] theorem lift_ofNat (n : ℕ) [n.AtLeastTwo] : - lift.{u} (no_index (OfNat.ofNat n : Cardinal.{v})) = OfNat.ofNat n := + lift.{u} (ofNat(n) : Cardinal.{v}) = OfNat.ofNat n := lift_natCast n @[simp] @@ -1302,7 +1301,7 @@ theorem lift_eq_nat_iff {a : Cardinal.{u}} {n : ℕ} : lift.{v} a = n ↔ a = n @[simp] theorem lift_eq_ofNat_iff {a : Cardinal.{u}} {n : ℕ} [n.AtLeastTwo] : - lift.{v} a = (no_index (OfNat.ofNat n)) ↔ a = OfNat.ofNat n := + lift.{v} a = ofNat(n) ↔ a = OfNat.ofNat n := lift_eq_nat_iff @[simp] @@ -1320,10 +1319,9 @@ theorem one_eq_lift_iff {a : Cardinal.{u}} : (1 : Cardinal) = lift.{v} a ↔ 1 = a := by simpa using nat_eq_lift_iff (n := 1) --- See note [no_index around OfNat.ofNat] @[simp] theorem ofNat_eq_lift_iff {a : Cardinal.{u}} {n : ℕ} [n.AtLeastTwo] : - (no_index (OfNat.ofNat n : Cardinal)) = lift.{v} a ↔ (OfNat.ofNat n : Cardinal) = a := + (ofNat(n) : Cardinal) = lift.{v} a ↔ (OfNat.ofNat n : Cardinal) = a := nat_eq_lift_iff @[simp] @@ -1335,10 +1333,9 @@ theorem lift_le_one_iff {a : Cardinal.{u}} : lift.{v} a ≤ 1 ↔ a ≤ 1 := by simpa using lift_le_nat_iff (n := 1) --- See note [no_index around OfNat.ofNat] @[simp] theorem lift_le_ofNat_iff {a : Cardinal.{u}} {n : ℕ} [n.AtLeastTwo] : - lift.{v} a ≤ (no_index (OfNat.ofNat n)) ↔ a ≤ OfNat.ofNat n := + lift.{v} a ≤ ofNat(n) ↔ a ≤ OfNat.ofNat n := lift_le_nat_iff @[simp] @@ -1350,27 +1347,24 @@ theorem one_le_lift_iff {a : Cardinal.{u}} : (1 : Cardinal) ≤ lift.{v} a ↔ 1 ≤ a := by simpa using nat_le_lift_iff (n := 1) --- See note [no_index around OfNat.ofNat] @[simp] theorem ofNat_le_lift_iff {a : Cardinal.{u}} {n : ℕ} [n.AtLeastTwo] : - (no_index (OfNat.ofNat n : Cardinal)) ≤ lift.{v} a ↔ (OfNat.ofNat n : Cardinal) ≤ a := + (ofNat(n) : Cardinal) ≤ lift.{v} a ↔ (OfNat.ofNat n : Cardinal) ≤ a := nat_le_lift_iff @[simp] theorem lift_lt_nat_iff {a : Cardinal.{u}} {n : ℕ} : lift.{v} a < n ↔ a < n := by rw [← lift_natCast.{v,u}, lift_lt] --- See note [no_index around OfNat.ofNat] @[simp] theorem lift_lt_ofNat_iff {a : Cardinal.{u}} {n : ℕ} [n.AtLeastTwo] : - lift.{v} a < (no_index (OfNat.ofNat n)) ↔ a < OfNat.ofNat n := + lift.{v} a < ofNat(n) ↔ a < OfNat.ofNat n := lift_lt_nat_iff @[simp] theorem nat_lt_lift_iff {n : ℕ} {a : Cardinal.{u}} : n < lift.{v} a ↔ n < a := by rw [← lift_natCast.{v,u}, lift_lt] --- See note [no_index around OfNat.ofNat] @[simp] theorem zero_lt_lift_iff {a : Cardinal.{u}} : (0 : Cardinal) < lift.{v} a ↔ 0 < a := by @@ -1381,10 +1375,9 @@ theorem one_lt_lift_iff {a : Cardinal.{u}} : (1 : Cardinal) < lift.{v} a ↔ 1 < a := by simpa using nat_lt_lift_iff (n := 1) --- See note [no_index around OfNat.ofNat] @[simp] theorem ofNat_lt_lift_iff {a : Cardinal.{u}} {n : ℕ} [n.AtLeastTwo] : - (no_index (OfNat.ofNat n : Cardinal)) < lift.{v} a ↔ (OfNat.ofNat n : Cardinal) < a := + (ofNat(n) : Cardinal) < lift.{v} a ↔ (OfNat.ofNat n : Cardinal) < a := nat_lt_lift_iff theorem lift_mk_fin (n : ℕ) : lift #(Fin n) = n := rfl @@ -1718,14 +1711,12 @@ theorem nat_mul_aleph0 {n : ℕ} (hn : n ≠ 0) : ↑n * ℵ₀ = ℵ₀ := @[simp] theorem aleph0_mul_nat {n : ℕ} (hn : n ≠ 0) : ℵ₀ * n = ℵ₀ := by rw [mul_comm, nat_mul_aleph0 hn] --- See note [no_index around OfNat.ofNat] @[simp] -theorem ofNat_mul_aleph0 {n : ℕ} [Nat.AtLeastTwo n] : no_index (OfNat.ofNat n) * ℵ₀ = ℵ₀ := +theorem ofNat_mul_aleph0 {n : ℕ} [Nat.AtLeastTwo n] : ofNat(n) * ℵ₀ = ℵ₀ := nat_mul_aleph0 (NeZero.ne n) --- See note [no_index around OfNat.ofNat] @[simp] -theorem aleph0_mul_ofNat {n : ℕ} [Nat.AtLeastTwo n] : ℵ₀ * no_index (OfNat.ofNat n) = ℵ₀ := +theorem aleph0_mul_ofNat {n : ℕ} [Nat.AtLeastTwo n] : ℵ₀ * ofNat(n) = ℵ₀ := aleph0_mul_nat (NeZero.ne n) @[simp] @@ -1740,14 +1731,12 @@ theorem aleph0_add_nat (n : ℕ) : ℵ₀ + n = ℵ₀ := @[simp] theorem nat_add_aleph0 (n : ℕ) : ↑n + ℵ₀ = ℵ₀ := by rw [add_comm, aleph0_add_nat] --- See note [no_index around OfNat.ofNat] @[simp] -theorem ofNat_add_aleph0 {n : ℕ} [Nat.AtLeastTwo n] : no_index (OfNat.ofNat n) + ℵ₀ = ℵ₀ := +theorem ofNat_add_aleph0 {n : ℕ} [Nat.AtLeastTwo n] : ofNat(n) + ℵ₀ = ℵ₀ := nat_add_aleph0 n --- See note [no_index around OfNat.ofNat] @[simp] -theorem aleph0_add_ofNat {n : ℕ} [Nat.AtLeastTwo n] : ℵ₀ + no_index (OfNat.ofNat n) = ℵ₀ := +theorem aleph0_add_ofNat {n : ℕ} [Nat.AtLeastTwo n] : ℵ₀ + ofNat(n) = ℵ₀ := aleph0_add_nat n theorem exists_nat_eq_of_le_nat {c : Cardinal} {n : ℕ} (h : c ≤ n) : ∃ m, m ≤ n ∧ c = m := by @@ -2204,4 +2193,4 @@ end Cardinal -- end Tactic -set_option linter.style.longFile 2400 +set_option linter.style.longFile 2300 diff --git a/Mathlib/SetTheory/Cardinal/Cofinality.lean b/Mathlib/SetTheory/Cardinal/Cofinality.lean index 5668b0eb5d566..c165757a42992 100644 --- a/Mathlib/SetTheory/Cardinal/Cofinality.lean +++ b/Mathlib/SetTheory/Cardinal/Cofinality.lean @@ -994,7 +994,7 @@ theorem le_range_of_union_finset_eq_top {α β : Type*} [Infinite β] (f : α exact infinite_univ by_contra h simp only [not_le] at h - let u : ∀ b, ∃ a, b ∈ f a := fun b => by simpa using (w.ge : _) (Set.mem_univ b) + let u : ∀ b, ∃ a, b ∈ f a := fun b => by simpa using (w.ge :) (Set.mem_univ b) let u' : β → range f := fun b => ⟨f (u b).choose, by simp⟩ have v' : ∀ a, u' ⁻¹' {⟨f a, by simp⟩} ≤ f a := by rintro a p m diff --git a/Mathlib/SetTheory/Cardinal/Continuum.lean b/Mathlib/SetTheory/Cardinal/Continuum.lean index c04a55eca2aa0..3f36674f9acb7 100644 --- a/Mathlib/SetTheory/Cardinal/Continuum.lean +++ b/Mathlib/SetTheory/Cardinal/Continuum.lean @@ -115,14 +115,12 @@ theorem nat_add_continuum (n : ℕ) : ↑n + 𝔠 = 𝔠 := theorem continuum_add_nat (n : ℕ) : 𝔠 + n = 𝔠 := (add_comm _ _).trans (nat_add_continuum n) --- See note [no_index around OfNat.ofNat] @[simp] -theorem ofNat_add_continuum {n : ℕ} [Nat.AtLeastTwo n] : no_index (OfNat.ofNat n) + 𝔠 = 𝔠 := +theorem ofNat_add_continuum {n : ℕ} [Nat.AtLeastTwo n] : ofNat(n) + 𝔠 = 𝔠 := nat_add_continuum n --- See note [no_index around OfNat.ofNat] @[simp] -theorem continuum_add_ofNat {n : ℕ} [Nat.AtLeastTwo n] : 𝔠 + no_index (OfNat.ofNat n) = 𝔠 := +theorem continuum_add_ofNat {n : ℕ} [Nat.AtLeastTwo n] : 𝔠 + ofNat(n) = 𝔠 := continuum_add_nat n /-! @@ -150,14 +148,12 @@ theorem nat_mul_continuum {n : ℕ} (hn : n ≠ 0) : ↑n * 𝔠 = 𝔠 := theorem continuum_mul_nat {n : ℕ} (hn : n ≠ 0) : 𝔠 * n = 𝔠 := (mul_comm _ _).trans (nat_mul_continuum hn) --- See note [no_index around OfNat.ofNat] @[simp] -theorem ofNat_mul_continuum {n : ℕ} [Nat.AtLeastTwo n] : no_index (OfNat.ofNat n) * 𝔠 = 𝔠 := +theorem ofNat_mul_continuum {n : ℕ} [Nat.AtLeastTwo n] : ofNat(n) * 𝔠 = 𝔠 := nat_mul_continuum (OfNat.ofNat_ne_zero n) --- See note [no_index around OfNat.ofNat] @[simp] -theorem continuum_mul_ofNat {n : ℕ} [Nat.AtLeastTwo n] : 𝔠 * no_index (OfNat.ofNat n) = 𝔠 := +theorem continuum_mul_ofNat {n : ℕ} [Nat.AtLeastTwo n] : 𝔠 * ofNat(n) = 𝔠 := continuum_mul_nat (OfNat.ofNat_ne_zero n) /-! diff --git a/Mathlib/SetTheory/Cardinal/ENat.lean b/Mathlib/SetTheory/Cardinal/ENat.lean index b444b8fb5cd1f..5ad5cee75b2b2 100644 --- a/Mathlib/SetTheory/Cardinal/ENat.lean +++ b/Mathlib/SetTheory/Cardinal/ENat.lean @@ -48,7 +48,7 @@ instance : Coe ENat Cardinal := ⟨Cardinal.ofENat⟩ @[simp, norm_cast] lemma ofENat_one : ofENat 1 = 1 := rfl @[simp, norm_cast] lemma ofENat_ofNat (n : ℕ) [n.AtLeastTwo] : - ((no_index (OfNat.ofNat n : ℕ∞)) : Cardinal) = OfNat.ofNat n := + ((ofNat(n) : ℕ∞) : Cardinal) = OfNat.ofNat n := rfl lemma ofENat_strictMono : StrictMono ofENat := @@ -67,14 +67,14 @@ lemma ofENat_lt_aleph0 {m : ℕ∞} : (m : Cardinal) < ℵ₀ ↔ m < ⊤ := @[simp] lemma ofENat_lt_nat {m : ℕ∞} {n : ℕ} : ofENat m < n ↔ m < n := by norm_cast @[simp] lemma ofENat_lt_ofNat {m : ℕ∞} {n : ℕ} [n.AtLeastTwo] : - ofENat m < no_index (OfNat.ofNat n) ↔ m < OfNat.ofNat n := ofENat_lt_nat + ofENat m < ofNat(n) ↔ m < OfNat.ofNat n := ofENat_lt_nat @[simp] lemma nat_lt_ofENat {m : ℕ} {n : ℕ∞} : (m : Cardinal) < n ↔ m < n := by norm_cast @[simp] lemma ofENat_pos {m : ℕ∞} : 0 < (m : Cardinal) ↔ 0 < m := by norm_cast @[simp] lemma one_lt_ofENat {m : ℕ∞} : 1 < (m : Cardinal) ↔ 1 < m := by norm_cast @[simp, norm_cast] lemma ofNat_lt_ofENat {m : ℕ} [m.AtLeastTwo] {n : ℕ∞} : - no_index (OfNat.ofNat m : Cardinal) < n ↔ OfNat.ofNat m < n := nat_lt_ofENat + (ofNat(m) : Cardinal) < n ↔ OfNat.ofNat m < n := nat_lt_ofENat lemma ofENat_mono : Monotone ofENat := ofENat_strictMono.monotone @@ -88,14 +88,14 @@ lemma ofENat_le_ofENat {m n : ℕ∞} : (m : Cardinal) ≤ n ↔ m ≤ n := ofEN @[simp] lemma ofENat_le_one {m : ℕ∞} : ofENat m ≤ 1 ↔ m ≤ 1 := by norm_cast @[simp] lemma ofENat_le_ofNat {m : ℕ∞} {n : ℕ} [n.AtLeastTwo] : - ofENat m ≤ no_index (OfNat.ofNat n) ↔ m ≤ OfNat.ofNat n := ofENat_le_nat + ofENat m ≤ ofNat(n) ↔ m ≤ OfNat.ofNat n := ofENat_le_nat @[simp] lemma nat_le_ofENat {m : ℕ} {n : ℕ∞} : (m : Cardinal) ≤ n ↔ m ≤ n := by norm_cast @[simp] lemma one_le_ofENat {n : ℕ∞} : 1 ≤ (n : Cardinal) ↔ 1 ≤ n := by norm_cast @[simp] lemma ofNat_le_ofENat {m : ℕ} [m.AtLeastTwo] {n : ℕ∞} : - no_index (OfNat.ofNat m : Cardinal) ≤ n ↔ OfNat.ofNat m ≤ n := nat_le_ofENat + (ofNat(m) : Cardinal) ≤ n ↔ OfNat.ofNat m ≤ n := nat_le_ofENat lemma ofENat_injective : Injective ofENat := ofENat_strictMono.injective @@ -112,10 +112,10 @@ lemma ofENat_inj {m n : ℕ∞} : (m : Cardinal) = n ↔ m = n := ofENat_injecti @[simp] lemma one_eq_ofENat {m : ℕ∞} : 1 = (m : Cardinal) ↔ m = 1 := by norm_cast; apply eq_comm @[simp] lemma ofENat_eq_ofNat {m : ℕ∞} {n : ℕ} [n.AtLeastTwo] : - (m : Cardinal) = no_index (OfNat.ofNat n) ↔ m = OfNat.ofNat n := ofENat_eq_nat + (m : Cardinal) = ofNat(n) ↔ m = OfNat.ofNat n := ofENat_eq_nat @[simp] lemma ofNat_eq_ofENat {m : ℕ} {n : ℕ∞} [m.AtLeastTwo] : - no_index (OfNat.ofNat m) = (n : Cardinal) ↔ OfNat.ofNat m = n := nat_eq_ofENat + ofNat(m) = (n : Cardinal) ↔ OfNat.ofNat m = n := nat_eq_ofENat @[simp, norm_cast] lemma lift_ofENat : ∀ m : ℕ∞, lift.{u, v} m = m | (m : ℕ) => lift_natCast m @@ -247,10 +247,10 @@ lemma toENat_nat (n : ℕ) : toENat n = n := map_natCast _ n @[simp] lemma toENat_eq_one {a : Cardinal} : toENat a = 1 ↔ a = 1 := toENat_eq_nat @[simp] lemma toENat_le_ofNat {a : Cardinal} {n : ℕ} [n.AtLeastTwo] : - toENat a ≤ no_index (OfNat.ofNat n) ↔ a ≤ OfNat.ofNat n := toENat_le_nat + toENat a ≤ ofNat(n) ↔ a ≤ OfNat.ofNat n := toENat_le_nat @[simp] lemma toENat_eq_ofNat {a : Cardinal} {n : ℕ} [n.AtLeastTwo] : - toENat a = no_index (OfNat.ofNat n) ↔ a = OfNat.ofNat n := toENat_eq_nat + toENat a = ofNat(n) ↔ a = OfNat.ofNat n := toENat_eq_nat @[simp] lemma toENat_eq_top {a : Cardinal} : toENat a = ⊤ ↔ ℵ₀ ≤ a := enat_gc.u_eq_top diff --git a/Mathlib/SetTheory/Cardinal/ToNat.lean b/Mathlib/SetTheory/Cardinal/ToNat.lean index 0763034be27cb..296a181875639 100644 --- a/Mathlib/SetTheory/Cardinal/ToNat.lean +++ b/Mathlib/SetTheory/Cardinal/ToNat.lean @@ -82,24 +82,12 @@ theorem toNat_lt_iff_lt_of_lt_aleph0 (hc : c < ℵ₀) (hd : d < ℵ₀) : theorem toNat_le_toNat (hcd : c ≤ d) (hd : d < ℵ₀) : toNat c ≤ toNat d := toNat_monotoneOn (hcd.trans_lt hd) hd hcd -@[deprecated toNat_le_toNat (since := "2024-02-15")] -theorem toNat_le_of_le_of_lt_aleph0 (hd : d < ℵ₀) (hcd : c ≤ d) : - toNat c ≤ toNat d := - toNat_le_toNat hcd hd - theorem toNat_lt_toNat (hcd : c < d) (hd : d < ℵ₀) : toNat c < toNat d := toNat_strictMonoOn (hcd.trans hd) hd hcd -@[deprecated toNat_lt_toNat (since := "2024-02-15")] -theorem toNat_lt_of_lt_of_lt_aleph0 (hd : d < ℵ₀) (hcd : c < d) : toNat c < toNat d := - toNat_lt_toNat hcd hd - -@[deprecated (since := "2024-02-15")] alias toNat_cast := toNat_natCast - --- See note [no_index around OfNat.ofNat] @[simp] theorem toNat_ofNat (n : ℕ) [n.AtLeastTwo] : - Cardinal.toNat (no_index (OfNat.ofNat n)) = OfNat.ofNat n := + Cardinal.toNat ofNat(n) = OfNat.ofNat n := toNat_natCast n /-- `toNat` has a right-inverse: coercion. -/ @@ -148,11 +136,6 @@ theorem toNat_congr {β : Type v} (e : α ≃ β) : toNat #α = toNat #β := by theorem toNat_mul (x y : Cardinal) : toNat (x * y) = toNat x * toNat y := map_mul toNat x y -@[deprecated map_prod (since := "2024-02-15")] -theorem toNat_finset_prod (s : Finset α) (f : α → Cardinal) : - toNat (∏ i ∈ s, f i) = ∏ i ∈ s, toNat (f i) := - map_prod toNat _ _ - @[simp] theorem toNat_add (hc : c < ℵ₀) (hd : d < ℵ₀) : toNat (c + d) = toNat c + toNat d := by lift c to ℕ using hc @@ -164,7 +147,4 @@ theorem toNat_lift_add_lift {a : Cardinal.{u}} {b : Cardinal.{v}} (ha : a < ℵ toNat (lift.{v} a + lift.{u} b) = toNat a + toNat b := by simp [*] -@[deprecated (since := "2024-02-15")] -alias toNat_add_of_lt_aleph0 := toNat_lift_add_lift - end Cardinal diff --git a/Mathlib/SetTheory/Game/Birthday.lean b/Mathlib/SetTheory/Game/Birthday.lean index c0da2ca0c18a8..09a00361b6f86 100644 --- a/Mathlib/SetTheory/Game/Birthday.lean +++ b/Mathlib/SetTheory/Game/Birthday.lean @@ -215,10 +215,9 @@ theorem birthday_natCast (n : ℕ) : birthday n = n := by rw [← toGame_natCast] exact birthday_ordinalToGame _ --- See note [no_index around OfNat.ofNat] @[simp] theorem birthday_ofNat (n : ℕ) [n.AtLeastTwo] : - birthday (no_index (OfNat.ofNat n)) = OfNat.ofNat n := + birthday ofNat(n) = OfNat.ofNat n := birthday_natCast n @[simp] diff --git a/Mathlib/SetTheory/Game/Short.lean b/Mathlib/SetTheory/Game/Short.lean index 0fd2f40701cd5..8ff13c9c88a07 100644 --- a/Mathlib/SetTheory/Game/Short.lean +++ b/Mathlib/SetTheory/Game/Short.lean @@ -209,7 +209,7 @@ instance shortNat : ∀ n : ℕ, Short n | 0 => PGame.short0 | n + 1 => @PGame.shortAdd _ _ (shortNat n) PGame.short1 -instance shortOfNat (n : ℕ) [Nat.AtLeastTwo n] : Short (no_index (OfNat.ofNat n)) := shortNat n +instance shortOfNat (n : ℕ) [Nat.AtLeastTwo n] : Short ofNat(n) := shortNat n instance shortBit0 (x : PGame.{u}) [Short x] : Short (x + x) := by infer_instance diff --git a/Mathlib/SetTheory/Ordinal/Arithmetic.lean b/Mathlib/SetTheory/Ordinal/Arithmetic.lean index c914bbd405ef9..058fe74f9c6cc 100644 --- a/Mathlib/SetTheory/Ordinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Ordinal/Arithmetic.lean @@ -2243,10 +2243,9 @@ theorem one_add_natCast (m : ℕ) : 1 + (m : Ordinal) = succ m := by @[deprecated "No deprecation message was provided." (since := "2024-04-17")] alias one_add_nat_cast := one_add_natCast --- See note [no_index around OfNat.ofNat] @[simp] theorem one_add_ofNat (m : ℕ) [m.AtLeastTwo] : - 1 + (no_index (OfNat.ofNat m : Ordinal)) = Order.succ (OfNat.ofNat m : Ordinal) := + 1 + (ofNat(m) : Ordinal) = Order.succ (OfNat.ofNat m : Ordinal) := one_add_natCast m @[simp, norm_cast] @@ -2334,10 +2333,9 @@ theorem lift_natCast : ∀ n : ℕ, lift.{u, v} n = n @[deprecated "No deprecation message was provided." (since := "2024-04-17")] alias lift_nat_cast := lift_natCast --- See note [no_index around OfNat.ofNat] @[simp] theorem lift_ofNat (n : ℕ) [n.AtLeastTwo] : - lift.{u, v} (no_index (OfNat.ofNat n)) = OfNat.ofNat n := + lift.{u, v} ofNat(n) = OfNat.ofNat n := lift_natCast n /-! ### Properties of ω -/ diff --git a/Mathlib/SetTheory/Ordinal/Basic.lean b/Mathlib/SetTheory/Ordinal/Basic.lean index 33a1d3e36a834..7f225262b4727 100644 --- a/Mathlib/SetTheory/Ordinal/Basic.lean +++ b/Mathlib/SetTheory/Ordinal/Basic.lean @@ -867,10 +867,9 @@ theorem type_sum_lex {α β : Type u} (r : α → α → Prop) (s : β → β theorem card_nat (n : ℕ) : card.{u} n = n := by induction n <;> [simp; simp only [card_add, card_one, Nat.cast_succ, *]] --- See note [no_index around OfNat.ofNat] @[simp] theorem card_ofNat (n : ℕ) [n.AtLeastTwo] : - card.{u} (no_index (OfNat.ofNat n)) = OfNat.ofNat n := + card.{u} ofNat(n) = OfNat.ofNat n := card_nat n instance instAddLeftMono : AddLeftMono Ordinal.{u} where @@ -1211,9 +1210,8 @@ theorem ord_nat (n : ℕ) : ord n = n := @[simp] theorem ord_one : ord 1 = 1 := by simpa using ord_nat 1 --- See note [no_index around OfNat.ofNat] @[simp] -theorem ord_ofNat (n : ℕ) [n.AtLeastTwo] : ord (no_index (OfNat.ofNat n)) = OfNat.ofNat n := +theorem ord_ofNat (n : ℕ) [n.AtLeastTwo] : ord ofNat(n) = OfNat.ofNat n := ord_nat n @[simp] @@ -1378,10 +1376,9 @@ theorem nat_le_card {o} {n : ℕ} : (n : Cardinal) ≤ card o ↔ (n : Ordinal) theorem one_le_card {o} : 1 ≤ card o ↔ 1 ≤ o := by simpa using nat_le_card (n := 1) --- See note [no_index around OfNat.ofNat] @[simp] theorem ofNat_le_card {o} {n : ℕ} [n.AtLeastTwo] : - (no_index (OfNat.ofNat n : Cardinal)) ≤ card o ↔ (OfNat.ofNat n : Ordinal) ≤ o := + (ofNat(n) : Cardinal) ≤ card o ↔ (OfNat.ofNat n : Ordinal) ≤ o := nat_le_card @[simp] @@ -1405,20 +1402,18 @@ theorem zero_lt_card {o} : 0 < card o ↔ 0 < o := by theorem one_lt_card {o} : 1 < card o ↔ 1 < o := by simpa using nat_lt_card (n := 1) --- See note [no_index around OfNat.ofNat] @[simp] theorem ofNat_lt_card {o} {n : ℕ} [n.AtLeastTwo] : - (no_index (OfNat.ofNat n : Cardinal)) < card o ↔ (OfNat.ofNat n : Ordinal) < o := + (ofNat(n) : Cardinal) < card o ↔ (OfNat.ofNat n : Ordinal) < o := nat_lt_card @[simp] theorem card_lt_nat {o} {n : ℕ} : card o < n ↔ o < n := lt_iff_lt_of_le_iff_le nat_le_card --- See note [no_index around OfNat.ofNat] @[simp] theorem card_lt_ofNat {o} {n : ℕ} [n.AtLeastTwo] : - card o < (no_index (OfNat.ofNat n)) ↔ o < OfNat.ofNat n := + card o < ofNat(n) ↔ o < OfNat.ofNat n := card_lt_nat @[simp] @@ -1429,10 +1424,9 @@ theorem card_le_nat {o} {n : ℕ} : card o ≤ n ↔ o ≤ n := theorem card_le_one {o} : card o ≤ 1 ↔ o ≤ 1 := by simpa using card_le_nat (n := 1) --- See note [no_index around OfNat.ofNat] @[simp] theorem card_le_ofNat {o} {n : ℕ} [n.AtLeastTwo] : - card o ≤ (no_index (OfNat.ofNat n)) ↔ o ≤ OfNat.ofNat n := + card o ≤ ofNat(n) ↔ o ≤ OfNat.ofNat n := card_le_nat @[simp] @@ -1457,10 +1451,9 @@ theorem lift_down' {a : Cardinal.{u}} {b : Ordinal.{max u v}} (h : card.{max u v} b ≤ Cardinal.lift.{v, u} a) : ∃ a', lift.{v, u} a' = b := mem_range_lift_of_card_le h --- See note [no_index around OfNat.ofNat] @[simp] theorem card_eq_ofNat {o} {n : ℕ} [n.AtLeastTwo] : - card o = (no_index (OfNat.ofNat n)) ↔ o = OfNat.ofNat n := + card o = ofNat(n) ↔ o = OfNat.ofNat n := card_eq_nat @[simp] diff --git a/Mathlib/Tactic/Hint.lean b/Mathlib/Tactic/Hint.lean index 8df08cd79a07e..a8a705aa76f8a 100644 --- a/Mathlib/Tactic/Hint.lean +++ b/Mathlib/Tactic/Hint.lean @@ -72,7 +72,7 @@ def suggestion (tac : TSyntax `tactic) (msgs : MessageLog := {}) : TacticM Sugge let msg? ← msgs.toList.findM? fun m => do pure <| m.severity == MessageSeverity.information && (← m.data.toString).startsWith "Try this: " let suggestion ← match msg? with - | some m => pure <| SuggestionText.string (((← m.data.toString).drop 10).takeWhile (· != '\n')) + | some m => pure <| SuggestionText.string ((← m.data.toString).drop 10) | none => pure <| SuggestionText.tsyntax tac return { suggestion, postInfo?, style? } diff --git a/Mathlib/Tactic/Linarith/Oracle/FourierMotzkin.lean b/Mathlib/Tactic/Linarith/Oracle/FourierMotzkin.lean index d1b3479c6c896..160faf10c2854 100644 --- a/Mathlib/Tactic/Linarith/Oracle/FourierMotzkin.lean +++ b/Mathlib/Tactic/Linarith/Oracle/FourierMotzkin.lean @@ -334,7 +334,7 @@ def CertificateOracle.fourierMotzkin : CertificateOracle where produceCertificate hyps maxVar := do let linarithData := mkLinarithData hyps maxVar let result ← - (ExceptT.run (StateT.run (do validate; elimAllVarsM : LinarithM Unit) linarithData) : _) + (ExceptT.run (StateT.run (do validate; elimAllVarsM : LinarithM Unit) linarithData) :) match result with | (Except.ok _) => failure | (Except.error contr) => return contr.src.flatten diff --git a/Mathlib/Tactic/Linter/HashCommandLinter.lean b/Mathlib/Tactic/Linter/HashCommandLinter.lean index 1f01c9b69a412..41832c0f4570b 100644 --- a/Mathlib/Tactic/Linter/HashCommandLinter.lean +++ b/Mathlib/Tactic/Linter/HashCommandLinter.lean @@ -65,12 +65,9 @@ This means that CI will eventually fail on `#`-commands, but does not stop it fr However, in order to avoid local clutter, when `warningAsError` is `false`, the linter logs a warning only for the `#`-commands that do not already emit a message. -/ def hashCommandLinter : Linter where run := withSetOptionIn' fun stx => do - let mod := (← getMainModule).components if Linter.getLinterValue linter.hashCommand (← getOptions) && - ((← get).messages.toList.isEmpty || warningAsError.get (← getOptions)) && - -- we check that the module is either not in `test` or, is `test.HashCommandLinter` - (mod.getD 0 default != `test || (mod == [`test, `HashCommandLinter])) - then + ((← get).messages.reportedPlusUnreported.isEmpty || warningAsError.get (← getOptions)) + then if let some sa := stx.getHead? then let a := sa.getAtomVal if (a.get ⟨0⟩ == '#' && ! allowed_commands.contains a) then diff --git a/Mathlib/Tactic/NoncommRing.lean b/Mathlib/Tactic/NoncommRing.lean index 7cced4a61594d..ad1410a6f8acb 100644 --- a/Mathlib/Tactic/NoncommRing.lean +++ b/Mathlib/Tactic/NoncommRing.lean @@ -23,9 +23,9 @@ namespace Mathlib.Tactic.NoncommRing section nat_lit_mul variable {R : Type*} [NonAssocSemiring R] (r : R) (n : ℕ) -lemma nat_lit_mul_eq_nsmul [n.AtLeastTwo] : no_index (OfNat.ofNat n) * r = OfNat.ofNat n • r := by +lemma nat_lit_mul_eq_nsmul [n.AtLeastTwo] : ofNat(n) * r = OfNat.ofNat n • r := by simp only [nsmul_eq_mul, Nat.cast_ofNat] -lemma mul_nat_lit_eq_nsmul [n.AtLeastTwo] : r * no_index (OfNat.ofNat n) = OfNat.ofNat n • r := by +lemma mul_nat_lit_eq_nsmul [n.AtLeastTwo] : r * ofNat(n) = OfNat.ofNat n • r := by simp only [nsmul_eq_mul', Nat.cast_ofNat] end nat_lit_mul diff --git a/Mathlib/Tactic/Ring/PNat.lean b/Mathlib/Tactic/Ring/PNat.lean index d5e2883e3d672..c37b0078f5085 100644 --- a/Mathlib/Tactic/Ring/PNat.lean +++ b/Mathlib/Tactic/Ring/PNat.lean @@ -25,6 +25,7 @@ instance {n} : CSLiftVal (no_index (OfNat.ofNat (n+1)) : ℕ+) (n + 1) := ⟨rfl instance {n h} : CSLiftVal (Nat.toPNat n h) n := ⟨rfl⟩ + instance {n} : CSLiftVal (Nat.succPNat n) (n + 1) := ⟨rfl⟩ instance {n} : CSLiftVal (Nat.toPNat' n) (n.pred + 1) := ⟨rfl⟩ diff --git a/Mathlib/Topology/Algebra/ConstMulAction.lean b/Mathlib/Topology/Algebra/ConstMulAction.lean index cd4bdb42717b6..29af329afbd4b 100644 --- a/Mathlib/Topology/Algebra/ConstMulAction.lean +++ b/Mathlib/Topology/Algebra/ConstMulAction.lean @@ -162,7 +162,7 @@ variable [Monoid M] [MulAction M α] [ContinuousConstSMul M α] @[to_additive] instance Units.continuousConstSMul : ContinuousConstSMul Mˣ α where - continuous_const_smul m := (continuous_const_smul (m : M) : _) + continuous_const_smul m := continuous_const_smul (m : M) @[to_additive] theorem smul_closure_subset (c : M) (s : Set α) : c • closure s ⊆ closure (c • s) := @@ -474,7 +474,7 @@ instance (priority := 100) t2Space_of_properlyDiscontinuousSMul_of_t2Space [T2Sp have f_op : IsOpenMap f := isOpenMap_quotient_mk'_mul rintro ⟨x₀⟩ ⟨y₀⟩ (hxy : f x₀ ≠ f y₀) show ∃ U ∈ 𝓝 (f x₀), ∃ V ∈ 𝓝 (f y₀), _ - have hγx₀y₀ : ∀ γ : Γ, γ • x₀ ≠ y₀ := not_exists.mp (mt Quotient.sound hxy.symm : _) + have hγx₀y₀ : ∀ γ : Γ, γ • x₀ ≠ y₀ := not_exists.mp (mt Quotient.sound hxy.symm :) obtain ⟨K₀, hK₀, K₀_in⟩ := exists_compact_mem_nhds x₀ obtain ⟨L₀, hL₀, L₀_in⟩ := exists_compact_mem_nhds y₀ let bad_Γ_set := { γ : Γ | (γ • ·) '' K₀ ∩ L₀ ≠ ∅ } diff --git a/Mathlib/Topology/Algebra/Group/Basic.lean b/Mathlib/Topology/Algebra/Group/Basic.lean index 003e76adb5395..28b8320d761a6 100644 --- a/Mathlib/Topology/Algebra/Group/Basic.lean +++ b/Mathlib/Topology/Algebra/Group/Basic.lean @@ -1008,7 +1008,7 @@ variable [TopologicalSpace α] {f g : α → G} {s : Set α} {x : α} @[to_additive (attr := continuity, fun_prop) sub] theorem Continuous.div' (hf : Continuous f) (hg : Continuous g) : Continuous fun x => f x / g x := - continuous_div'.comp (hf.prod_mk hg : _) + continuous_div'.comp (hf.prod_mk hg :) @[to_additive (attr := continuity) continuous_sub_left] lemma continuous_div_left' (a : G) : Continuous (a / ·) := continuous_const.div' continuous_id diff --git a/Mathlib/Topology/Algebra/InfiniteSum/Constructions.lean b/Mathlib/Topology/Algebra/InfiniteSum/Constructions.lean index 4386fe90f6bd1..81e5cd9c09b6a 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/Constructions.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/Constructions.lean @@ -256,14 +256,14 @@ open MulOpposite variable [AddCommMonoid α] [TopologicalSpace α] {f : β → α} {a : α} theorem HasSum.op (hf : HasSum f a) : HasSum (fun a ↦ op (f a)) (op a) := - (hf.map (@opAddEquiv α _) continuous_op : _) + (hf.map (@opAddEquiv α _) continuous_op :) theorem Summable.op (hf : Summable f) : Summable (op ∘ f) := hf.hasSum.op.summable theorem HasSum.unop {f : β → αᵐᵒᵖ} {a : αᵐᵒᵖ} (hf : HasSum f a) : HasSum (fun a ↦ unop (f a)) (unop a) := - (hf.map (@opAddEquiv α _).symm continuous_unop : _) + (hf.map (@opAddEquiv α _).symm continuous_unop :) theorem Summable.unop {f : β → αᵐᵒᵖ} (hf : Summable f) : Summable (unop ∘ f) := hf.hasSum.unop.summable diff --git a/Mathlib/Topology/Algebra/InfiniteSum/NatInt.lean b/Mathlib/Topology/Algebra/InfiniteSum/NatInt.lean index 688be889d23d8..662314f0bb691 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/NatInt.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/NatInt.lean @@ -354,7 +354,6 @@ lemma HasProd.of_nat_of_neg_add_one {f : ℤ → M} exact (Nat.cast_injective.hasProd_range_iff.mpr hf₁).mul_isCompl this (hi₂.hasProd_range_iff.mpr hf₂) -@[deprecated (since := "2024-03-04")] alias HasSum.nonneg_add_neg := HasSum.of_nat_of_neg_add_one @[to_additive Summable.of_nat_of_neg_add_one] lemma Multipliable.of_nat_of_neg_add_one {f : ℤ → M} @@ -437,9 +436,6 @@ theorem HasProd.nat_mul_neg {f : ℤ → M} (hf : HasProd f m) : simp only [u1, u2, Nat.cast_inj, imp_self, implies_true, forall_const, prod_image, neg_inj] _ = ∏ b ∈ v', (f b * f (-b)) := prod_mul_distrib.symm⟩ -@[deprecated HasSum.nat_add_neg (since := "2024-03-04")] -alias HasSum.sum_nat_of_sum_int := HasSum.nat_add_neg - @[to_additive] theorem Multipliable.nat_mul_neg {f : ℤ → M} (hf : Multipliable f) : Multipliable fun n : ℕ ↦ f n * f (-n) := @@ -456,9 +452,6 @@ theorem HasProd.of_add_one_of_neg_add_one {f : ℤ → M} HasProd f (m * f 0 * m') := HasProd.of_nat_of_neg_add_one (mul_comm _ m ▸ HasProd.zero_mul hf₁) hf₂ -@[deprecated HasSum.of_add_one_of_neg_add_one (since := "2024-03-04")] -alias HasSum.pos_add_zero_add_neg := HasSum.of_add_one_of_neg_add_one - @[to_additive Summable.of_add_one_of_neg_add_one] lemma Multipliable.of_add_one_of_neg_add_one {f : ℤ → M} (hf₁ : Multipliable fun n : ℕ ↦ f (n + 1)) (hf₂ : Multipliable fun n : ℕ ↦ f (-(n + 1))) : @@ -490,9 +483,6 @@ lemma Multipliable.of_nat_of_neg {f : ℤ → G} (hf₁ : Multipliable fun n : (hf₂ : Multipliable fun n : ℕ ↦ f (-n)) : Multipliable f := (hf₁.hasProd.of_nat_of_neg hf₂.hasProd).multipliable -@[deprecated Summable.of_nat_of_neg (since := "2024-03-04")] -alias summable_int_of_summable_nat := Summable.of_nat_of_neg - @[to_additive] lemma tprod_of_nat_of_neg [T2Space G] {f : ℤ → G} (hf₁ : Multipliable fun n : ℕ ↦ f n) (hf₂ : Multipliable fun n : ℕ ↦ f (-n)) : @@ -524,3 +514,16 @@ lemma multipliable_int_iff_multipliable_nat_and_neg {f : ℤ → G} : end UniformGroup end Int + +section pnat + +@[to_additive] +theorem pnat_multipliable_iff_multipliable_succ {α : Type*} [TopologicalSpace α] [CommMonoid α] + {f : ℕ → α} : Multipliable (fun x : ℕ+ => f x) ↔ Multipliable fun x : ℕ => f (x + 1) := + Equiv.pnatEquivNat.symm.multipliable_iff.symm + +@[to_additive] +theorem tprod_pnat_eq_tprod_succ {α : Type*} [TopologicalSpace α] [CommMonoid α] (f : ℕ → α) : + ∏' n : ℕ+, f n = ∏' n, f (n + 1) := (Equiv.pnatEquivNat.symm.tprod_eq _).symm + +end pnat diff --git a/Mathlib/Topology/Algebra/Module/LinearMap.lean b/Mathlib/Topology/Algebra/Module/LinearMap.lean index 48a01db20692e..c11c4e9d10354 100644 --- a/Mathlib/Topology/Algebra/Module/LinearMap.lean +++ b/Mathlib/Topology/Algebra/Module/LinearMap.lean @@ -505,7 +505,7 @@ theorem natCast_apply [ContinuousAdd M₁] (n : ℕ) (m : M₁) : (↑n : M₁ @[simp] theorem ofNat_apply [ContinuousAdd M₁] (n : ℕ) [n.AtLeastTwo] (m : M₁) : - ((no_index (OfNat.ofNat n) : M₁ →L[R₁] M₁)) m = OfNat.ofNat n • m := + (ofNat(n) : M₁ →L[R₁] M₁) m = OfNat.ofNat n • m := rfl section ApplyAction diff --git a/Mathlib/Topology/Algebra/Module/Multilinear/Basic.lean b/Mathlib/Topology/Algebra/Module/Multilinear/Basic.lean index 65348675488a9..442b6637f5b9f 100644 --- a/Mathlib/Topology/Algebra/Module/Multilinear/Basic.lean +++ b/Mathlib/Topology/Algebra/Module/Multilinear/Basic.lean @@ -299,6 +299,55 @@ theorem _root_.ContinuousLinearMap.compContinuousMultilinearMap_coe (g : M₂ ext m rfl +/-- `ContinuousMultilinearMap.prod` as an `Equiv`. -/ +@[simps apply symm_apply_fst symm_apply_snd, simps (config := .lemmasOnly) symm_apply] +def prodEquiv : + (ContinuousMultilinearMap R M₁ M₂ × ContinuousMultilinearMap R M₁ M₃) ≃ + ContinuousMultilinearMap R M₁ (M₂ × M₃) where + toFun f := f.1.prod f.2 + invFun f := ((ContinuousLinearMap.fst _ _ _).compContinuousMultilinearMap f, + (ContinuousLinearMap.snd _ _ _).compContinuousMultilinearMap f) + left_inv _ := rfl + right_inv _ := rfl + +theorem prod_ext_iff {f g : ContinuousMultilinearMap R M₁ (M₂ × M₃)} : + f = g ↔ (ContinuousLinearMap.fst _ _ _).compContinuousMultilinearMap f = + (ContinuousLinearMap.fst _ _ _).compContinuousMultilinearMap g ∧ + (ContinuousLinearMap.snd _ _ _).compContinuousMultilinearMap f = + (ContinuousLinearMap.snd _ _ _).compContinuousMultilinearMap g := by + rw [← Prod.mk.inj_iff, ← prodEquiv_symm_apply, ← prodEquiv_symm_apply, Equiv.apply_eq_iff_eq] + +@[ext] +theorem prod_ext {f g : ContinuousMultilinearMap R M₁ (M₂ × M₃)} + (h₁ : (ContinuousLinearMap.fst _ _ _).compContinuousMultilinearMap f = + (ContinuousLinearMap.fst _ _ _).compContinuousMultilinearMap g) + (h₂ : (ContinuousLinearMap.snd _ _ _).compContinuousMultilinearMap f = + (ContinuousLinearMap.snd _ _ _).compContinuousMultilinearMap g) : f = g := + prod_ext_iff.mpr ⟨h₁, h₂⟩ + +theorem eq_prod_iff {f : ContinuousMultilinearMap R M₁ (M₂ × M₃)} + {g : ContinuousMultilinearMap R M₁ M₂} {h : ContinuousMultilinearMap R M₁ M₃} : + f = g.prod h ↔ (ContinuousLinearMap.fst _ _ _).compContinuousMultilinearMap f = g ∧ + (ContinuousLinearMap.snd _ _ _).compContinuousMultilinearMap f = h := + prod_ext_iff + +theorem add_prod_add [ContinuousAdd M₂] [ContinuousAdd M₃] + (f₁ f₂ : ContinuousMultilinearMap R M₁ M₂) (g₁ g₂ : ContinuousMultilinearMap R M₁ M₃) : + (f₁ + f₂).prod (g₁ + g₂) = f₁.prod g₁ + f₂.prod g₂ := + rfl + +theorem smul_prod_smul {S : Type*} [Monoid S] [DistribMulAction S M₂] [DistribMulAction S M₃] + [ContinuousConstSMul S M₂] [SMulCommClass R S M₂] + [ContinuousConstSMul S M₃] [SMulCommClass R S M₃] + (c : S) (f : ContinuousMultilinearMap R M₁ M₂) (g : ContinuousMultilinearMap R M₁ M₃) : + (c • f).prod (c • g) = c • f.prod g := + rfl + +@[simp] +theorem zero_prod_zero : + (0 : ContinuousMultilinearMap R M₁ M₂).prod (0 : ContinuousMultilinearMap R M₁ M₃) = 0 := + rfl + /-- `ContinuousMultilinearMap.pi` as an `Equiv`. -/ @[simps] def piEquiv {ι' : Type*} {M' : ι' → Type*} [∀ i, AddCommMonoid (M' i)] @@ -306,12 +355,8 @@ def piEquiv {ι' : Type*} {M' : ι' → Type*} [∀ i, AddCommMonoid (M' i)] (∀ i, ContinuousMultilinearMap R M₁ (M' i)) ≃ ContinuousMultilinearMap R M₁ (∀ i, M' i) where toFun := ContinuousMultilinearMap.pi invFun f i := (ContinuousLinearMap.proj i : _ →L[R] M' i).compContinuousMultilinearMap f - left_inv f := by - ext - rfl - right_inv f := by - ext - rfl + left_inv _ := rfl + right_inv _ := rfl /-- An equivalence of the index set defines an equivalence between the spaces of continuous multilinear maps. This is the forward map of this equivalence. -/ @@ -453,6 +498,16 @@ instance : AddCommGroup (ContinuousMultilinearMap R M₁ M₂) := toMultilinearMap_injective.addCommGroup _ rfl (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ _ => rfl +theorem neg_prod_neg [AddCommGroup M₃] [Module R M₃] [TopologicalSpace M₃] [TopologicalAddGroup M₃] + (f : ContinuousMultilinearMap R M₁ M₂) (g : ContinuousMultilinearMap R M₁ M₃) : + (-f).prod (-g) = - f.prod g := + rfl + +theorem sub_prod_sub [AddCommGroup M₃] [Module R M₃] [TopologicalSpace M₃] [TopologicalAddGroup M₃] + (f₁ f₂ : ContinuousMultilinearMap R M₁ M₂) (g₁ g₂ : ContinuousMultilinearMap R M₁ M₃) : + (f₁ - f₂).prod (g₁ - g₂) = f₁.prod g₁ - f₂.prod g₂ := + rfl + end TopologicalAddGroup end Ring diff --git a/Mathlib/Topology/Algebra/Monoid.lean b/Mathlib/Topology/Algebra/Monoid.lean index ccd647dbad6f8..f136b5fd39568 100644 --- a/Mathlib/Topology/Algebra/Monoid.lean +++ b/Mathlib/Topology/Algebra/Monoid.lean @@ -89,7 +89,7 @@ theorem ContinuousMul.induced {α : Type*} {β : Type*} {F : Type*} [FunLike F @[to_additive (attr := continuity, fun_prop)] theorem Continuous.mul {f g : X → M} (hf : Continuous f) (hg : Continuous g) : Continuous fun x => f x * g x := - continuous_mul.comp (hf.prod_mk hg : _) + continuous_mul.comp (hf.prod_mk hg :) @[to_additive (attr := continuity)] theorem continuous_mul_left (a : M) : Continuous fun b : M => a * b := @@ -102,7 +102,7 @@ theorem continuous_mul_right (a : M) : Continuous fun b : M => b * a := @[to_additive (attr := fun_prop)] theorem ContinuousOn.mul {f g : X → M} {s : Set X} (hf : ContinuousOn f s) (hg : ContinuousOn g s) : ContinuousOn (fun x => f x * g x) s := - (continuous_mul.comp_continuousOn (hf.prod hg) : _) + (continuous_mul.comp_continuousOn (hf.prod hg) :) @[to_additive] theorem tendsto_mul {a b : M} : Tendsto (fun p : M × M => p.fst * p.snd) (𝓝 (a, b)) (𝓝 (a * b)) := diff --git a/Mathlib/Topology/Algebra/Order/Field.lean b/Mathlib/Topology/Algebra/Order/Field.lean index 8bb48fa820d3e..4faff225f1e90 100644 --- a/Mathlib/Topology/Algebra/Order/Field.lean +++ b/Mathlib/Topology/Algebra/Order/Field.lean @@ -116,11 +116,20 @@ lemma inv_atTop₀ : (atTop : Filter 𝕜)⁻¹ = 𝓝[>] 0 := (((atTop_basis_Ioi' (0 : 𝕜)).map _).comp_surjective inv_surjective).eq_of_same_basis <| (nhdsGT_basis _).congr (by simp) fun a ha ↦ by simp [inv_Ioi₀ (inv_pos.2 ha)] +@[simp] +lemma inv_atBot₀ : (atBot : Filter 𝕜)⁻¹ = 𝓝[<] 0 := + (((atBot_basis_Iio' (0 : 𝕜)).map _).comp_surjective inv_surjective).eq_of_same_basis <| + (nhdsLT_basis _).congr (by simp) fun a ha ↦ by simp [inv_Iio₀ (inv_neg''.2 ha)] + @[simp] lemma inv_nhdsGT_zero : (𝓝[>] (0 : 𝕜))⁻¹ = atTop := by rw [← inv_atTop₀, inv_inv] @[deprecated (since := "2024-12-22")] alias inv_nhdsWithin_Ioi_zero := inv_nhdsGT_zero +@[simp] +lemma inv_nhdsLT_zero : (𝓝[<] (0 : 𝕜))⁻¹ = atBot := by + rw [← inv_atBot₀, inv_inv] + /-- The function `x ↦ x⁻¹` tends to `+∞` on the right of `0`. -/ theorem tendsto_inv_nhdsGT_zero : Tendsto (fun x : 𝕜 => x⁻¹) (𝓝[>] (0 : 𝕜)) atTop := inv_nhdsGT_zero.le @@ -138,24 +147,50 @@ alias tendsto_inv_atTop_zero' := tendsto_inv_atTop_nhdsGT_zero theorem tendsto_inv_atTop_zero : Tendsto (fun r : 𝕜 => r⁻¹) atTop (𝓝 0) := tendsto_inv_atTop_nhdsGT_zero.mono_right inf_le_left +/-- The function `x ↦ x⁻¹` tends to `-∞` on the left of `0`. -/ +theorem tendsto_inv_zero_atBot : Tendsto (fun x : 𝕜 => x⁻¹) (𝓝[<] (0 : 𝕜)) atBot := + inv_nhdsLT_zero.le + +/-- The function `r ↦ r⁻¹` tends to `0` on the left as `r → -∞`. -/ +theorem tendsto_inv_atBot_zero' : Tendsto (fun r : 𝕜 => r⁻¹) atBot (𝓝[<] (0 : 𝕜)) := + inv_atBot₀.le + +theorem tendsto_inv_atBot_zero : Tendsto (fun r : 𝕜 => r⁻¹) atBot (𝓝 0) := + tendsto_inv_atBot_zero'.mono_right inf_le_left + theorem Filter.Tendsto.div_atTop {a : 𝕜} (h : Tendsto f l (𝓝 a)) (hg : Tendsto g l atTop) : Tendsto (fun x => f x / g x) l (𝓝 0) := by simp only [div_eq_mul_inv] exact mul_zero a ▸ h.mul (tendsto_inv_atTop_zero.comp hg) +theorem Filter.Tendsto.div_atBot {a : 𝕜} (h : Tendsto f l (𝓝 a)) (hg : Tendsto g l atBot) : + Tendsto (fun x => f x / g x) l (𝓝 0) := by + simp only [div_eq_mul_inv] + exact mul_zero a ▸ h.mul (tendsto_inv_atBot_zero.comp hg) + lemma Filter.Tendsto.const_div_atTop (hg : Tendsto g l atTop) (r : 𝕜) : Tendsto (fun n ↦ r / g n) l (𝓝 0) := tendsto_const_nhds.div_atTop hg +lemma Filter.Tendsto.const_div_atBot (hg : Tendsto g l atBot) (r : 𝕜) : + Tendsto (fun n ↦ r / g n) l (𝓝 0) := + tendsto_const_nhds.div_atBot hg + theorem Filter.Tendsto.inv_tendsto_atTop (h : Tendsto f l atTop) : Tendsto f⁻¹ l (𝓝 0) := tendsto_inv_atTop_zero.comp h +theorem Filter.Tendsto.inv_tendsto_atBot (h : Tendsto f l atBot) : Tendsto f⁻¹ l (𝓝 0) := + tendsto_inv_atBot_zero.comp h + theorem Filter.Tendsto.inv_tendsto_nhdsGT_zero (h : Tendsto f l (𝓝[>] 0)) : Tendsto f⁻¹ l atTop := tendsto_inv_nhdsGT_zero.comp h @[deprecated (since := "2024-12-22")] alias Filter.Tendsto.inv_tendsto_zero := Filter.Tendsto.inv_tendsto_nhdsGT_zero +theorem Filter.Tendsto.inv_tendsto_nhdsLT_zero (h : Tendsto f l (𝓝[<] 0)) : Tendsto f⁻¹ l atBot := + tendsto_inv_zero_atBot.comp h + /-- If `g` tends to zero and there exists a constant `C : 𝕜` such that eventually `|f x| ≤ C`, then the product `f * g` tends to zero. -/ theorem bdd_le_mul_tendsto_zero' {f g : α → 𝕜} (C : 𝕜) (hf : ∀ᶠ x in l, |f x| ≤ C) diff --git a/Mathlib/Topology/Algebra/SeparationQuotient/Basic.lean b/Mathlib/Topology/Algebra/SeparationQuotient/Basic.lean index 14e8443fb52c5..536e68588f287 100644 --- a/Mathlib/Topology/Algebra/SeparationQuotient/Basic.lean +++ b/Mathlib/Topology/Algebra/SeparationQuotient/Basic.lean @@ -268,7 +268,7 @@ theorem mk_natCast [NatCast R] (n : ℕ) : mk (n : R) = n := rfl @[simp] theorem mk_ofNat [NatCast R] (n : ℕ) [n.AtLeastTwo] : - mk (no_index (OfNat.ofNat n) : R) = OfNat.ofNat n := + mk (ofNat(n) : R) = OfNat.ofNat n := rfl instance instIntCast [IntCast R] : IntCast (SeparationQuotient R) where diff --git a/Mathlib/Topology/Algebra/UniformField.lean b/Mathlib/Topology/Algebra/UniformField.lean index 8c8c66bc9a593..a3f84dc3d6317 100644 --- a/Mathlib/Topology/Algebra/UniformField.lean +++ b/Mathlib/Topology/Algebra/UniformField.lean @@ -124,7 +124,7 @@ theorem mul_hatInv_cancel {x : hat K} (x_ne : x ≠ 0) : x * hatInv x = 1 := by letI : TopologicalSpace (hat K × hat K) := instTopologicalSpaceProd have : ContinuousAt (fun y : hat K => ((y, hatInv y) : hat K × hat K)) x := continuous_id.continuousAt.prod (continuous_hatInv x_ne) - exact (_root_.continuous_mul.continuousAt.comp this : _) + exact (_root_.continuous_mul.continuousAt.comp this :) have clo : x ∈ closure (c '' {0}ᶜ) := by have := isDenseInducing_coe.dense x rw [← image_univ, show (univ : Set K) = {0} ∪ {0}ᶜ from (union_compl_self _).symm, diff --git a/Mathlib/Topology/Algebra/UniformGroup/Basic.lean b/Mathlib/Topology/Algebra/UniformGroup/Basic.lean index 4baeea0bda780..f02e105e37e93 100644 --- a/Mathlib/Topology/Algebra/UniformGroup/Basic.lean +++ b/Mathlib/Topology/Algebra/UniformGroup/Basic.lean @@ -279,7 +279,7 @@ private theorem extend_Z_bilin_aux (x₀ : α) (y₁ : δ) : ∃ U₂ ∈ comap have := Tendsto.prod_mk (tendsto_sub_comap_self de x₀) (tendsto_const_nhds : Tendsto (fun _ : β × β => y₁) (comap ee <| 𝓝 (x₀, x₀)) (𝓝 y₁)) rw [nhds_prod_eq, prod_comap_comap_eq, ← nhds_prod_eq] - exact (this : _) + exact (this :) have lim2 : Tendsto (fun p : β × δ => φ p.1 p.2) (𝓝 (0, y₁)) (𝓝 0) := by simpa using hφ.tendsto (0, y₁) have lim := lim2.comp lim1 diff --git a/Mathlib/Topology/Algebra/UniformGroup/Defs.lean b/Mathlib/Topology/Algebra/UniformGroup/Defs.lean index 5b6a3aeeb9ff2..2733cc8e8d6b6 100644 --- a/Mathlib/Topology/Algebra/UniformGroup/Defs.lean +++ b/Mathlib/Topology/Algebra/UniformGroup/Defs.lean @@ -479,7 +479,7 @@ private theorem extend_Z_bilin_aux (x₀ : α) (y₁ : δ) : ∃ U₂ ∈ comap have := Tendsto.prod_mk (tendsto_sub_comap_self de x₀) (tendsto_const_nhds : Tendsto (fun _ : β × β => y₁) (comap ee <| 𝓝 (x₀, x₀)) (𝓝 y₁)) rw [nhds_prod_eq, prod_comap_comap_eq, ← nhds_prod_eq] - exact (this : _) + exact (this :) have lim2 : Tendsto (fun p : β × δ => φ p.1 p.2) (𝓝 (0, y₁)) (𝓝 0) := by simpa using hφ.tendsto (0, y₁) have lim := lim2.comp lim1 diff --git a/Mathlib/Topology/Algebra/UniformRing.lean b/Mathlib/Topology/Algebra/UniformRing.lean index a52f63d8fafe8..5453277e0629f 100644 --- a/Mathlib/Topology/Algebra/UniformRing.lean +++ b/Mathlib/Topology/Algebra/UniformRing.lean @@ -233,16 +233,6 @@ theorem inseparableSetoid_ring (α) [CommRing α] [TopologicalSpace α] [Topolog Setoid.ext fun x y => addGroup_inseparable_iff.trans <| .trans (by rfl) (Submodule.quotientRel_def _).symm -@[deprecated (since := "2024-03-09")] -alias ring_sep_rel := inseparableSetoid_ring - --- Equality of types is evil -@[deprecated UniformSpace.inseparableSetoid_ring (since := "2024-02-16")] -theorem ring_sep_quot (α : Type u) [r : CommRing α] [TopologicalSpace α] [TopologicalRing α] : - SeparationQuotient α = (α ⧸ (⊥ : Ideal α).closure) := by - rw [SeparationQuotient, @inseparableSetoid_ring α r] - rfl - /-- Given a topological ring `α` equipped with a uniform structure that makes subtraction uniformly continuous, get an homeomorphism between the separated quotient of `α` and the quotient ring corresponding to the closure of zero. -/ diff --git a/Mathlib/Topology/Bases.lean b/Mathlib/Topology/Bases.lean index 691e389d9e045..db461531876af 100644 --- a/Mathlib/Topology/Bases.lean +++ b/Mathlib/Topology/Bases.lean @@ -526,15 +526,9 @@ theorem isSeparable_range [TopologicalSpace β] [SeparableSpace α] {f : α → theorem IsSeparable.of_subtype (s : Set α) [SeparableSpace s] : IsSeparable s := by simpa using isSeparable_range (continuous_subtype_val (p := (· ∈ s))) -@[deprecated (since := "2024-02-05")] -alias isSeparable_of_separableSpace_subtype := IsSeparable.of_subtype - theorem IsSeparable.of_separableSpace [h : SeparableSpace α] (s : Set α) : IsSeparable s := IsSeparable.mono (isSeparable_univ_iff.2 h) (subset_univ _) -@[deprecated (since := "2024-02-05")] -alias isSeparable_of_separableSpace := IsSeparable.of_separableSpace - end TopologicalSpace open TopologicalSpace diff --git a/Mathlib/Topology/Basic.lean b/Mathlib/Topology/Basic.lean index e6f284b71d21c..04de611bd1b5b 100644 --- a/Mathlib/Topology/Basic.lean +++ b/Mathlib/Topology/Basic.lean @@ -1665,7 +1665,7 @@ However, lemmas with this conclusion are not nice to use in practice because continuous_add.comp (continuous_id.prod_mk continuous_id) ``` The second is a valid proof, which is accepted if you write it as - `continuous_add.comp (continuous_id.prod_mk continuous_id : _)` + `continuous_add.comp (continuous_id.prod_mk continuous_id :)` 2. If the operation has more than 2 arguments, they are impractical to use, because in your application the arguments in the domain might be in a different order or associated differently. diff --git a/Mathlib/Topology/Category/Profinite/AsLimit.lean b/Mathlib/Topology/Category/Profinite/AsLimit.lean index 72fb66b943303..a417cb6fdd284 100644 --- a/Mathlib/Topology/Category/Profinite/AsLimit.lean +++ b/Mathlib/Topology/Category/Profinite/AsLimit.lean @@ -45,7 +45,7 @@ def fintypeDiagram : DiscreteQuotient X ⥤ FintypeCat where map f := DiscreteQuotient.ofLE f.le -- Porting note: `map_comp` used to be proved by default by `aesop_cat`. -- once `aesop_cat` can prove this again, remove the entire `map_comp` here. - map_comp _ _ := by ext; aesop_cat + map_comp _ _ := by ext; simp /-- An abbreviation for `X.fintypeDiagram ⋙ FintypeCat.toProfinite`. -/ abbrev diagram : DiscreteQuotient X ⥤ Profinite := diff --git a/Mathlib/Topology/Category/TopCat/Limits/Basic.lean b/Mathlib/Topology/Category/TopCat/Limits/Basic.lean index 98981cae50ee9..99c6519501cba 100644 --- a/Mathlib/Topology/Category/TopCat/Limits/Basic.lean +++ b/Mathlib/Topology/Category/TopCat/Limits/Basic.lean @@ -100,7 +100,7 @@ def limitConeInfiIsLimit (F : J ⥤ TopCat.{max v u}) : IsLimit (limitConeInfi.{ continuous_iff_coinduced_le.mpr (le_iInf fun j => coinduced_le_iff_le_induced.mp <| - (continuous_iff_coinduced_le.mp (s.π.app j).continuous : _)) + (continuous_iff_coinduced_le.mp (s.π.app j).continuous :)) · rfl instance topCat_hasLimitsOfSize : HasLimitsOfSize.{w, v} TopCat.{max v u} where @@ -159,7 +159,7 @@ def colimitCoconeIsColimit (F : J ⥤ TopCat.{max v u}) : IsColimit (colimitCoco continuous_iff_le_induced.mpr (iSup_le fun j => coinduced_le_iff_le_induced.mp <| - (continuous_iff_coinduced_le.mp (s.ι.app j).continuous : _)) + (continuous_iff_coinduced_le.mp (s.ι.app j).continuous :)) · rfl instance topCat_hasColimitsOfSize : HasColimitsOfSize.{w,v} TopCat.{max v u} where diff --git a/Mathlib/Topology/Category/TopCat/Limits/Products.lean b/Mathlib/Topology/Category/TopCat/Limits/Products.lean index 6512a54157e40..6a1dbd61c0562 100644 --- a/Mathlib/Topology/Category/TopCat/Limits/Products.lean +++ b/Mathlib/Topology/Category/TopCat/Limits/Products.lean @@ -58,12 +58,12 @@ theorem piIsoPi_inv_π {ι : Type v} (α : ι → TopCat.{max v u}) (i : ι) : (piIsoPi α).inv ≫ Pi.π α i = piπ α i := by simp [piIsoPi] theorem piIsoPi_inv_π_apply {ι : Type v} (α : ι → TopCat.{max v u}) (i : ι) (x : ∀ i, α i) : - (Pi.π α i : _) ((piIsoPi α).inv x) = x i := + (Pi.π α i :) ((piIsoPi α).inv x) = x i := ConcreteCategory.congr_hom (piIsoPi_inv_π α i) x -- Porting note: needing the type ascription on `∏ᶜ α : TopCat.{max v u}` is unfortunate. theorem piIsoPi_hom_apply {ι : Type v} (α : ι → TopCat.{max v u}) (i : ι) - (x : (∏ᶜ α : TopCat.{max v u})) : (piIsoPi α).hom x i = (Pi.π α i : _) x := by + (x : (∏ᶜ α : TopCat.{max v u})) : (piIsoPi α).hom x i = (Pi.π α i :) x := by have := piIsoPi_inv_π α i rw [Iso.inv_comp_eq] at this exact ConcreteCategory.congr_hom this x @@ -106,11 +106,11 @@ theorem sigmaIsoSigma_hom_ι {ι : Type v} (α : ι → TopCat.{max v u}) (i : Sigma.ι α i ≫ (sigmaIsoSigma α).hom = sigmaι α i := by simp [sigmaIsoSigma] theorem sigmaIsoSigma_hom_ι_apply {ι : Type v} (α : ι → TopCat.{max v u}) (i : ι) (x : α i) : - (sigmaIsoSigma α).hom ((Sigma.ι α i : _) x) = Sigma.mk i x := + (sigmaIsoSigma α).hom ((Sigma.ι α i :) x) = Sigma.mk i x := ConcreteCategory.congr_hom (sigmaIsoSigma_hom_ι α i) x theorem sigmaIsoSigma_inv_apply {ι : Type v} (α : ι → TopCat.{max v u}) (i : ι) (x : α i) : - (sigmaIsoSigma α).inv ⟨i, x⟩ = (Sigma.ι α i : _) x := by + (sigmaIsoSigma α).inv ⟨i, x⟩ = (Sigma.ι α i :) x := by rw [← sigmaIsoSigma_hom_ι_apply, ← comp_app, ← comp_app, Iso.hom_inv_id, Category.comp_id] @@ -277,7 +277,7 @@ def binaryCofanIsColimit (X Y : TopCat.{u}) : IsColimit (TopCat.binaryCofan X Y) rfl · intro s m h₁ h₂ ext (x | x) - exacts [(ConcreteCategory.congr_hom h₁ x : _), (ConcreteCategory.congr_hom h₂ x : _)] + exacts [(ConcreteCategory.congr_hom h₁ x :), (ConcreteCategory.congr_hom h₂ x :)] theorem binaryCofan_isColimit_iff {X Y : TopCat} (c : BinaryCofan X Y) : Nonempty (IsColimit c) ↔ diff --git a/Mathlib/Topology/Category/TopCat/Limits/Pullbacks.lean b/Mathlib/Topology/Category/TopCat/Limits/Pullbacks.lean index f7dfb76f0da23..1bc507cd752fd 100644 --- a/Mathlib/Topology/Category/TopCat/Limits/Pullbacks.lean +++ b/Mathlib/Topology/Category/TopCat/Limits/Pullbacks.lean @@ -452,14 +452,14 @@ theorem colimit_topology (F : J ⥤ TopCat.{max v u}) : (colimit F).str = ⨆ j, (F.obj j).str.coinduced (colimit.ι F j) := coinduced_of_isColimit _ (colimit.isColimit F) -theorem colimit_isOpen_iff (F : J ⥤ TopCat.{max v u}) (U : Set ((colimit F : _) : Type max v u)) : +theorem colimit_isOpen_iff (F : J ⥤ TopCat.{max v u}) (U : Set ((colimit F :) : Type max v u)) : IsOpen U ↔ ∀ j, IsOpen (colimit.ι F j ⁻¹' U) := by dsimp [topologicalSpace_coe] conv_lhs => rw [colimit_topology F] exact isOpen_iSup_iff theorem coequalizer_isOpen_iff (F : WalkingParallelPair ⥤ TopCat.{u}) - (U : Set ((colimit F : _) : Type u)) : + (U : Set ((colimit F :) : Type u)) : IsOpen U ↔ IsOpen (colimit.ι F WalkingParallelPair.one ⁻¹' U) := by rw [colimit_isOpen_iff] constructor diff --git a/Mathlib/Topology/Category/TopCat/Opens.lean b/Mathlib/Topology/Category/TopCat/Opens.lean index 3a0f0d9137b1e..4e4930958fd0e 100644 --- a/Mathlib/Topology/Category/TopCat/Opens.lean +++ b/Mathlib/Topology/Category/TopCat/Opens.lean @@ -91,7 +91,7 @@ theorem infLELeft_apply_mk (U V : Opens X) (x) (m) : @[simp] theorem leSupr_apply_mk {ι : Type*} (U : ι → Opens X) (i : ι) (x) (m) : - (leSupr U i) ⟨x, m⟩ = ⟨x, (le_iSup U i : _) m⟩ := + (leSupr U i) ⟨x, m⟩ = ⟨x, (le_iSup U i :) m⟩ := rfl /-- The functor from open sets in `X` to `TopCat`, diff --git a/Mathlib/Topology/CompactOpen.lean b/Mathlib/Topology/CompactOpen.lean index 49f2b425934f9..3f9c290325a74 100644 --- a/Mathlib/Topology/CompactOpen.lean +++ b/Mathlib/Topology/CompactOpen.lean @@ -279,17 +279,11 @@ theorem compactOpen_eq_iInf_induced : convert isOpen_induced (isOpen_setOf_mapsTo (isCompact_iff_isCompact_univ.1 hK) hU) simp [mapsTo_univ_iff, Subtype.forall, MapsTo] -@[deprecated (since := "2024-03-05")] -alias compactOpen_eq_sInf_induced := compactOpen_eq_iInf_induced - theorem nhds_compactOpen_eq_iInf_nhds_induced (f : C(X, Y)) : 𝓝 f = ⨅ (s) (_ : IsCompact s), (𝓝 (f.restrict s)).comap (ContinuousMap.restrict s) := by rw [compactOpen_eq_iInf_induced] simp only [nhds_iInf, nhds_induced] -@[deprecated (since := "2024-03-05")] -alias nhds_compactOpen_eq_sInf_nhds_induced := nhds_compactOpen_eq_iInf_nhds_induced - theorem tendsto_compactOpen_restrict {ι : Type*} {l : Filter ι} {F : ι → C(X, Y)} {f : C(X, Y)} (hFf : Filter.Tendsto F l (𝓝 f)) (s : Set X) : Tendsto (fun i => (F i).restrict s) l (𝓝 (f.restrict s)) := @@ -371,15 +365,6 @@ def curry (f : C(X × Y, Z)) : C(X, C(Y, Z)) where theorem curry_apply (f : C(X × Y, Z)) (a : X) (b : Y) : f.curry a b = f (a, b) := rfl -/-- Auxiliary definition, see `ContinuousMap.curry` and `Homeomorph.curry`. -/ -@[deprecated ContinuousMap.curry (since := "2024-03-05")] -def curry' (f : C(X × Y, Z)) (a : X) : C(Y, Z) := curry f a - -set_option linter.deprecated false in -/-- If a map `α × β → γ` is continuous, then its curried form `α → C(β, γ)` is continuous. -/ -@[deprecated ContinuousMap.curry (since := "2024-03-05")] -theorem continuous_curry' (f : C(X × Y, Z)) : Continuous (curry' f) := (curry f).continuous - /-- To show continuity of a map `α → C(β, γ)`, it suffices to show that its uncurried form `α × β → γ` is continuous. -/ theorem continuous_of_continuous_uncurry (f : X → C(Y, Z)) diff --git a/Mathlib/Topology/Compactness/Compact.lean b/Mathlib/Topology/Compactness/Compact.lean index 15e1b303bf533..3482d2379b813 100644 --- a/Mathlib/Topology/Compactness/Compact.lean +++ b/Mathlib/Topology/Compactness/Compact.lean @@ -283,10 +283,6 @@ theorem IsCompact.nonempty_iInter_of_directed_nonempty_isCompact_isClosed rcases htd i₀ i with ⟨j, hji₀, hji⟩ exact (htn j).mono (subset_inter hji₀ hji) -@[deprecated (since := "2024-02-28")] -alias IsCompact.nonempty_iInter_of_directed_nonempty_compact_closed := - IsCompact.nonempty_iInter_of_directed_nonempty_isCompact_isClosed - /-- Cantor's intersection theorem for `sInter`: the intersection of a directed family of nonempty compact closed sets is nonempty. -/ theorem IsCompact.nonempty_sInter_of_directed_nonempty_isCompact_isClosed @@ -307,10 +303,6 @@ theorem IsCompact.nonempty_iInter_of_sequence_nonempty_isCompact_isClosed (t : have htc : ∀ i, IsCompact (t i) := fun i => ht0.of_isClosed_subset (htcl i) (this i) IsCompact.nonempty_iInter_of_directed_nonempty_isCompact_isClosed t htd htn htc htcl -@[deprecated (since := "2024-02-28")] -alias IsCompact.nonempty_iInter_of_sequence_nonempty_compact_closed := - IsCompact.nonempty_iInter_of_sequence_nonempty_isCompact_isClosed - /-- For every open cover of a compact set, there exists a finite subcover. -/ theorem IsCompact.elim_finite_subcover_image {b : Set ι} {c : ι → Set X} (hs : IsCompact s) (hc₁ : ∀ i ∈ b, IsOpen (c i)) (hc₂ : s ⊆ ⋃ i ∈ b, c i) : @@ -640,14 +632,6 @@ theorem mem_coclosedCompact_iff : closure_minimal (compl_subset_comm.2 hst) htcl · exact ⟨closure sᶜ, ⟨isClosed_closure, h⟩, compl_subset_comm.2 subset_closure⟩ -@[deprecated mem_coclosedCompact_iff (since := "2024-02-16")] -theorem mem_coclosedCompact : s ∈ coclosedCompact X ↔ ∃ t, IsClosed t ∧ IsCompact t ∧ tᶜ ⊆ s := by - simp only [hasBasis_coclosedCompact.mem_iff, and_assoc] - -@[deprecated mem_coclosedCompact_iff (since := "2024-02-16")] -theorem mem_coclosed_compact' : s ∈ coclosedCompact X ↔ ∃ t, IsClosed t ∧ IsCompact t ∧ sᶜ ⊆ t := by - simp only [hasBasis_coclosedCompact.mem_iff, compl_subset_comm, and_assoc] - /-- Complement of a set belongs to `coclosedCompact` if and only if its closure is compact. -/ theorem compl_mem_coclosedCompact : sᶜ ∈ coclosedCompact X ↔ IsCompact (closure s) := by rw [mem_coclosedCompact_iff, compl_compl] diff --git a/Mathlib/Topology/Constructions.lean b/Mathlib/Topology/Constructions.lean index ee4d5a39eb471..5e5987a52ca6d 100644 --- a/Mathlib/Topology/Constructions.lean +++ b/Mathlib/Topology/Constructions.lean @@ -467,8 +467,6 @@ theorem Continuous.uncurry_right {f : X → Y → Z} (y : Y) (h : Continuous (un Continuous fun a => f a y := h.comp (Continuous.Prod.mk_left _) -@[deprecated (since := "2024-03-09")] alias continuous_uncurry_left := Continuous.uncurry_left -@[deprecated (since := "2024-03-09")] alias continuous_uncurry_right := Continuous.uncurry_right theorem continuous_curry {g : X × Y → Z} (x : X) (h : Continuous g) : Continuous (curry g x) := Continuous.uncurry_left x h diff --git a/Mathlib/Topology/ContinuousMap/Algebra.lean b/Mathlib/Topology/ContinuousMap/Algebra.lean index 7b2984309d3c0..5d745326a17a6 100644 --- a/Mathlib/Topology/ContinuousMap/Algebra.lean +++ b/Mathlib/Topology/ContinuousMap/Algebra.lean @@ -50,7 +50,7 @@ variable [TopologicalSpace α] [TopologicalSpace β] [TopologicalSpace γ] @[to_additive] instance instMul [Mul β] [ContinuousMul β] : Mul C(α, β) := - ⟨fun f g => ⟨f * g, continuous_mul.comp (f.continuous.prod_mk g.continuous : _)⟩⟩ + ⟨fun f g => ⟨f * g, continuous_mul.comp (f.continuous.prod_mk g.continuous :)⟩⟩ @[to_additive (attr := norm_cast, simp)] theorem coe_mul [Mul β] [ContinuousMul β] (f g : C(α, β)) : ⇑(f * g) = f * g := diff --git a/Mathlib/Topology/ContinuousMap/Bounded/Basic.lean b/Mathlib/Topology/ContinuousMap/Bounded/Basic.lean index c634039f9359d..6a26326badb7b 100644 --- a/Mathlib/Topology/ContinuousMap/Bounded/Basic.lean +++ b/Mathlib/Topology/ContinuousMap/Bounded/Basic.lean @@ -519,7 +519,7 @@ theorem arzela_ascoli₁ [CompactSpace β] (A : Set (α →ᵇ β)) (closed : Is · exact (hU x').2.2 _ hx' _ (hU x').1 hf · exact (hU x').2.2 _ hx' _ (hU x').1 hg · have F_f_g : F (f x') = F (g x') := - (congr_arg (fun f : tα → tβ => (f ⟨x', x'tα⟩ : β)) f_eq_g : _) + (congr_arg (fun f : tα → tβ => (f ⟨x', x'tα⟩ : β)) f_eq_g :) calc dist (f x') (g x') ≤ dist (f x') (F (f x')) + dist (g x') (F (f x')) := dist_triangle_right _ _ _ @@ -1414,8 +1414,6 @@ instance instLattice : Lattice (α →ᵇ β) := DFunLike.coe_injective.lattice @[simp, norm_cast] lemma coe_posPart (f : α →ᵇ β) : ⇑f⁺ = (⇑f)⁺ := rfl @[simp, norm_cast] lemma coe_negPart (f : α →ᵇ β) : ⇑f⁻ = (⇑f)⁻ := rfl -@[deprecated (since := "2024-02-21")] alias coeFn_sup := coe_sup -@[deprecated (since := "2024-02-21")] alias coeFn_abs := coe_abs instance instNormedLatticeAddCommGroup : NormedLatticeAddCommGroup (α →ᵇ β) := { instSeminormedAddCommGroup with diff --git a/Mathlib/Topology/ContinuousMap/StoneWeierstrass.lean b/Mathlib/Topology/ContinuousMap/StoneWeierstrass.lean index e62e6489c8c9e..9e88d03fbe726 100644 --- a/Mathlib/Topology/ContinuousMap/StoneWeierstrass.lean +++ b/Mathlib/Topology/ContinuousMap/StoneWeierstrass.lean @@ -67,11 +67,8 @@ theorem polynomial_comp_attachBound (A : Subalgebra ℝ C(X, ℝ)) (f : A) (g : (g.toContinuousMapOn (Set.Icc (-‖f‖) ‖f‖)).comp (f : C(X, ℝ)).attachBound = Polynomial.aeval f g := by ext - simp only [ContinuousMap.coe_comp, Function.comp_apply, ContinuousMap.attachBound_apply_coe, - Polynomial.toContinuousMapOn_apply, Polynomial.aeval_subalgebra_coe, - Polynomial.aeval_continuousMap_apply, Polynomial.toContinuousMap_apply] - -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 - erw [ContinuousMap.attachBound_apply_coe] + simp only [Polynomial.aeval_subalgebra_coe, Polynomial.aeval_continuousMap_apply] + simp /-- Given a continuous function `f` in a subalgebra of `C(X, ℝ)`, postcomposing by a polynomial gives another function in `A`. diff --git a/Mathlib/Topology/ContinuousMap/ZeroAtInfty.lean b/Mathlib/Topology/ContinuousMap/ZeroAtInfty.lean index eecb451c02ae7..8756075a38d8e 100644 --- a/Mathlib/Topology/ContinuousMap/ZeroAtInfty.lean +++ b/Mathlib/Topology/ContinuousMap/ZeroAtInfty.lean @@ -430,7 +430,6 @@ theorem isClosed_range_toBCF : IsClosed (range (toBCF : C₀(α, β) → α → _ < ε := by simpa [add_halves ε] using add_lt_add_right (mem_ball.1 hg) (ε / 2) exact ⟨⟨f.toContinuousMap, this⟩, rfl⟩ -@[deprecated (since := "2024-03-17")] alias closed_range_toBCF := isClosed_range_toBCF /-- Continuous functions vanishing at infinity taking values in a complete space form a complete space. -/ @@ -467,7 +466,7 @@ variable [SeminormedAddCommGroup β] {𝕜 : Type*} [NormedField 𝕜] [NormedSp theorem norm_toBCF_eq_norm {f : C₀(α, β)} : ‖f.toBCF‖ = ‖f‖ := rfl -instance : NormedSpace 𝕜 C₀(α, β) where norm_smul_le k f := (norm_smul_le k f.toBCF : _) +instance : NormedSpace 𝕜 C₀(α, β) where norm_smul_le k f := (norm_smul_le k f.toBCF :) end NormedSpace @@ -535,7 +534,7 @@ section NormedStar variable [NormedAddCommGroup β] [StarAddMonoid β] [NormedStarGroup β] instance instNormedStarGroup : NormedStarGroup C₀(α, β) where - norm_star f := (norm_star f.toBCF : _) + norm_star f := (norm_star f.toBCF :) end NormedStar diff --git a/Mathlib/Topology/ContinuousOn.lean b/Mathlib/Topology/ContinuousOn.lean index 745a728ee1127..9147bdcc6b1f6 100644 --- a/Mathlib/Topology/ContinuousOn.lean +++ b/Mathlib/Topology/ContinuousOn.lean @@ -215,7 +215,7 @@ theorem union_mem_nhds_of_mem_nhdsWithin {b : α} {L : Set α} (hL : L ∈ nhdsWithin b I₁) {R : Set α} (hR : R ∈ nhdsWithin b I₂) : L ∪ R ∈ nhds b := by rw [← nhdsWithin_univ b, h, nhdsWithin_union] - exact ⟨mem_of_superset hL (by aesop), mem_of_superset hR (by aesop)⟩ + exact ⟨mem_of_superset hL (by simp), mem_of_superset hR (by simp)⟩ /-- Writing a punctured neighborhood filter as a sup of left and right filters. -/ diff --git a/Mathlib/Topology/Filter.lean b/Mathlib/Topology/Filter.lean index 675480013b398..6cdd957ccc15e 100644 --- a/Mathlib/Topology/Filter.lean +++ b/Mathlib/Topology/Filter.lean @@ -119,16 +119,16 @@ theorem nhds_pure (x : α) : 𝓝 (pure x : Filter α) = 𝓟 {⊥, pure x} := b rw [← principal_singleton, nhds_principal, principal_singleton, Iic_pure] @[simp] -theorem nhds_iInf (f : ι → Filter α) : 𝓝 (⨅ i, f i) = ⨅ i, 𝓝 (f i) := by +protected theorem nhds_iInf (f : ι → Filter α) : 𝓝 (⨅ i, f i) = ⨅ i, 𝓝 (f i) := by simp only [nhds_eq] apply lift'_iInf_of_map_univ <;> simp @[simp] -theorem nhds_inf (l₁ l₂ : Filter α) : 𝓝 (l₁ ⊓ l₂) = 𝓝 l₁ ⊓ 𝓝 l₂ := by - simpa only [iInf_bool_eq] using nhds_iInf fun b => cond b l₁ l₂ +protected theorem nhds_inf (l₁ l₂ : Filter α) : 𝓝 (l₁ ⊓ l₂) = 𝓝 l₁ ⊓ 𝓝 l₂ := by + simpa only [iInf_bool_eq] using Filter.nhds_iInf fun b => cond b l₁ l₂ theorem monotone_nhds : Monotone (𝓝 : Filter α → Filter (Filter α)) := - Monotone.of_map_inf nhds_inf + Monotone.of_map_inf Filter.nhds_inf theorem sInter_nhds (l : Filter α) : ⋂₀ { s | s ∈ 𝓝 l } = Iic l := by simp_rw [nhds_eq, Function.comp_def, sInter_lift'_sets monotone_principal.Iic, Iic, @@ -164,7 +164,7 @@ instance : T0Space (Filter α) := (specializes_iff_le.1 h.symm.specializes)⟩ theorem nhds_atTop [Preorder α] : 𝓝 atTop = ⨅ x : α, 𝓟 (Iic (𝓟 (Ici x))) := by - simp only [atTop, nhds_iInf, nhds_principal] + simp only [atTop, Filter.nhds_iInf, nhds_principal] protected theorem tendsto_nhds_atTop_iff [Preorder β] {l : Filter α} {f : α → Filter β} : Tendsto f l (𝓝 atTop) ↔ ∀ y, ∀ᶠ a in l, Ici y ∈ f a := by diff --git a/Mathlib/Topology/GDelta/Basic.lean b/Mathlib/Topology/GDelta/Basic.lean index 60e2975ce2246..250e1184d4569 100644 --- a/Mathlib/Topology/GDelta/Basic.lean +++ b/Mathlib/Topology/GDelta/Basic.lean @@ -66,25 +66,21 @@ theorem IsOpen.isGδ {s : Set X} (h : IsOpen s) : IsGδ s := protected theorem IsGδ.empty : IsGδ (∅ : Set X) := isOpen_empty.isGδ -@[deprecated (since := "2024-02-15")] alias isGδ_empty := IsGδ.empty @[simp] protected theorem IsGδ.univ : IsGδ (univ : Set X) := isOpen_univ.isGδ -@[deprecated (since := "2024-02-15")] alias isGδ_univ := IsGδ.univ 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 [forall_mem_image], hI.image _, by rw [sInter_image]⟩ -@[deprecated (since := "2024-02-15")] alias isGδ_biInter_of_isOpen := IsGδ.biInter_of_isOpen theorem IsGδ.iInter_of_isOpen [Countable ι'] {f : ι' → Set X} (hf : ∀ i, IsOpen (f i)) : IsGδ (⋂ i, f i) := ⟨range f, by rwa [forall_mem_range], countable_range _, by rw [sInter_range]⟩ -@[deprecated (since := "2024-02-15")] alias isGδ_iInter_of_isOpen := IsGδ.iInter_of_isOpen lemma isGδ_iff_eq_iInter_nat {s : Set X} : IsGδ s ↔ ∃ (f : ℕ → Set X), (∀ n, IsOpen (f n)) ∧ s = ⋂ n, f n := by @@ -115,13 +111,11 @@ theorem IsGδ.biInter {s : Set ι} (hs : s.Countable) {t : ∀ i ∈ s, Set X} haveI := hs.to_subtype exact .iInter fun x => ht x x.2 -@[deprecated (since := "2024-02-15")] alias isGδ_biInter := IsGδ.biInter /-- A countable intersection of Gδ sets is a Gδ set. -/ theorem IsGδ.sInter {S : Set (Set X)} (h : ∀ s ∈ S, IsGδ s) (hS : S.Countable) : IsGδ (⋂₀ S) := by simpa only [sInter_eq_biInter] using IsGδ.biInter hS h -@[deprecated (since := "2024-02-15")] alias isGδ_sInter := IsGδ.sInter theorem IsGδ.inter {s t : Set X} (hs : IsGδ s) (ht : IsGδ t) : IsGδ (s ∩ t) := by rw [inter_eq_iInter] @@ -150,9 +144,6 @@ theorem IsGδ.biUnion {s : Set ι} (hs : s.Finite) {f : ι → Set X} (h : ∀ i rw [← sUnion_image] exact .sUnion (hs.image _) (forall_mem_image.2 h) -@[deprecated (since := "2024-02-15")] -alias isGδ_biUnion := IsGδ.biUnion - /-- The union of finitely many Gδ sets is a Gδ set, bounded indexed union version. -/ theorem IsGδ.iUnion [Finite ι'] {f : ι' → Set X} (h : ∀ i, IsGδ (f i)) : IsGδ (⋃ i, f i) := .sUnion (finite_range _) <| forall_mem_range.2 h diff --git a/Mathlib/Topology/GDelta/UniformSpace.lean b/Mathlib/Topology/GDelta/UniformSpace.lean index 3aab2f9dd00c0..735d3d2bd0269 100644 --- a/Mathlib/Topology/GDelta/UniformSpace.lean +++ b/Mathlib/Topology/GDelta/UniformSpace.lean @@ -50,6 +50,5 @@ theorem IsGδ.setOf_continuousAt [UniformSpace Y] [IsCountablyGenerated (𝓤 Y) rintro ⟨s, ⟨hsx, hso⟩, hsU⟩ filter_upwards [IsOpen.mem_nhds hso hsx] with _ hy using ⟨s, ⟨hy, hso⟩, hsU⟩ -@[deprecated (since := "2024-02-15")] alias isGδ_setOf_continuousAt := IsGδ.setOf_continuousAt end ContinuousAt diff --git a/Mathlib/Topology/Germ.lean b/Mathlib/Topology/Germ.lean index 35d47bcfa9b8a..acee607839565 100644 --- a/Mathlib/Topology/Germ.lean +++ b/Mathlib/Topology/Germ.lean @@ -17,7 +17,7 @@ with respect to the neighbourhood filter `𝓝 x`. * `Filter.Germ.value φ f`: value associated to the germ `φ` at a point `x`, w.r.t. the neighbourhood filter at `x`. This is the common value of all representatives of `φ` at `x`. * `Filter.Germ.valueOrderRingHom` and friends: the map `Germ (𝓝 x) E → E` is a -monoid homomorphism, 𝕜-module homomorphism, ring homomorphism, monotone ring homomorphism +monoid homomorphism, 𝕜-linear map, ring homomorphism, monotone ring homomorphism * `RestrictGermPredicate`: given a predicate on germs `P : Π x : X, germ (𝓝 x) Y → Prop` and `A : set X`, build a new predicate on germs `restrictGermPredicate P A` such that diff --git a/Mathlib/Topology/Gluing.lean b/Mathlib/Topology/Gluing.lean index 9af7fc85de01f..2b05f77e1de94 100644 --- a/Mathlib/Topology/Gluing.lean +++ b/Mathlib/Topology/Gluing.lean @@ -195,7 +195,7 @@ theorem ι_eq_iff_rel (i j : D.J) (x : D.U i) (y : D.U j) : show _ = Sigma.mk j y from ConcreteCategory.congr_hom (sigmaIsoSigma.{_, u} D.U).inv_hom_id _] change InvImage D.Rel (sigmaIsoSigma.{_, u} D.U).hom _ _ rw [← (InvImage.equivalence _ _ D.rel_equiv).eqvGen_iff] - refine Relation.EqvGen.mono ?_ (D.eqvGen_of_π_eq h : _) + refine Relation.EqvGen.mono ?_ (D.eqvGen_of_π_eq h :) rintro _ _ ⟨x⟩ obtain ⟨⟨⟨i, j⟩, y⟩, rfl⟩ := (ConcreteCategory.bijective_of_isIso (sigmaIsoSigma.{u, u} _).inv).2 x diff --git a/Mathlib/Topology/Instances/Discrete.lean b/Mathlib/Topology/Instances/Discrete.lean index b9af3e7060b0e..4acd770481069 100644 --- a/Mathlib/Topology/Instances/Discrete.lean +++ b/Mathlib/Topology/Instances/Discrete.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ import Mathlib.Topology.Order.Basic +import Mathlib.Order.SuccPred.LinearLocallyFinite /-! # Instances related to the discrete topology @@ -36,11 +37,6 @@ instance (priority := 100) DiscreteTopology.secondCountableTopology_of_countable secondCountableTopology_of_countable_cover (singletons_open_iff_discrete.mpr hd) (iUnion_of_singleton α) -@[deprecated DiscreteTopology.secondCountableTopology_of_countable (since := "2024-03-11")] -theorem DiscreteTopology.secondCountableTopology_of_encodable {α : Type*} - [TopologicalSpace α] [DiscreteTopology α] [Countable α] : SecondCountableTopology α := - DiscreteTopology.secondCountableTopology_of_countable - theorem LinearOrder.bot_topologicalSpace_eq_generateFrom {α} [LinearOrder α] [PredOrder α] [SuccOrder α] : (⊥ : TopologicalSpace α) = generateFrom { s | ∃ a, s = Ioi a ∨ s = Iio a } := by let _ := Preorder.topology α @@ -63,3 +59,9 @@ alias discreteTopology_iff_orderTopology_of_pred_succ' := instance OrderTopology.of_discreteTopology [LinearOrder α] [PredOrder α] [SuccOrder α] [DiscreteTopology α] : OrderTopology α := discreteTopology_iff_orderTopology_of_pred_succ.mp ‹_› + +instance OrderTopology.of_linearLocallyFinite + [LinearOrder α] [LocallyFiniteOrder α] [DiscreteTopology α] : OrderTopology α := + haveI := LinearLocallyFiniteOrder.succOrder α + haveI := LinearLocallyFiniteOrder.predOrder α + inferInstance diff --git a/Mathlib/Topology/Instances/ENNReal.lean b/Mathlib/Topology/Instances/ENNReal.lean index fe2ae506c5631..93f04f375c31f 100644 --- a/Mathlib/Topology/Instances/ENNReal.lean +++ b/Mathlib/Topology/Instances/ENNReal.lean @@ -288,6 +288,35 @@ protected theorem tendsto_atTop_zero [Nonempty β] [SemilatticeSup β] {f : β Tendsto f atTop (𝓝 0) ↔ ∀ ε > 0, ∃ N, ∀ n ≥ N, f n ≤ ε := .trans (atTop_basis.tendsto_iff nhds_zero_basis_Iic) (by simp only [true_and]; rfl) +theorem tendsto_atTop_zero_iff_le_of_antitone {β : Type*} [Nonempty β] [SemilatticeSup β] + {f : β → ℝ≥0∞} (hf : Antitone f) : + Filter.Tendsto f Filter.atTop (𝓝 0) ↔ ∀ ε, 0 < ε → ∃ n : β, f n ≤ ε := by + rw [ENNReal.tendsto_atTop_zero] + refine ⟨fun h ↦ fun ε hε ↦ ?_, fun h ↦ fun ε hε ↦ ?_⟩ + · obtain ⟨n, hn⟩ := h ε hε + exact ⟨n, hn n le_rfl⟩ + · obtain ⟨n, hn⟩ := h ε hε + exact ⟨n, fun m hm ↦ (hf hm).trans hn⟩ + +theorem tendsto_atTop_zero_iff_lt_of_antitone {β : Type*} [Nonempty β] [SemilatticeSup β] + {f : β → ℝ≥0∞} (hf : Antitone f) : + Filter.Tendsto f Filter.atTop (𝓝 0) ↔ ∀ ε, 0 < ε → ∃ n : β, f n < ε := by + rw [ENNReal.tendsto_atTop_zero_iff_le_of_antitone hf] + constructor <;> intro h ε hε + · obtain ⟨n, hn⟩ := h (min 1 (ε / 2)) + (lt_min_iff.mpr ⟨zero_lt_one, (ENNReal.div_pos_iff.mpr ⟨ne_of_gt hε, ENNReal.two_ne_top⟩)⟩) + · refine ⟨n, hn.trans_lt ?_⟩ + by_cases hε_top : ε = ∞ + · rw [hε_top] + exact (min_le_left _ _).trans_lt ENNReal.one_lt_top + refine (min_le_right _ _).trans_lt ?_ + rw [ENNReal.div_lt_iff (Or.inr hε.ne') (Or.inr hε_top)] + conv_lhs => rw [← mul_one ε] + rw [ENNReal.mul_lt_mul_left hε.ne' hε_top] + norm_num + · obtain ⟨n, hn⟩ := h ε hε + exact ⟨n, hn.le⟩ + theorem tendsto_sub : ∀ {a b : ℝ≥0∞}, (a ≠ ∞ ∨ b ≠ ∞) → Tendsto (fun p : ℝ≥0∞ × ℝ≥0∞ => p.1 - p.2) (𝓝 (a, b)) (𝓝 (a - b)) | ∞, ∞, h => by simp only [ne_eq, not_true_eq_false, or_self] at h @@ -1176,7 +1205,7 @@ theorem continuous_edist : Continuous fun p : α × α => edist p.1 p.2 := by @[continuity, fun_prop] theorem Continuous.edist [TopologicalSpace β] {f g : β → α} (hf : Continuous f) (hg : Continuous g) : Continuous fun b => edist (f b) (g b) := - continuous_edist.comp (hf.prod_mk hg : _) + continuous_edist.comp (hf.prod_mk hg :) theorem Filter.Tendsto.edist {f g : β → α} {x : Filter β} {a b : α} (hf : Tendsto f x (𝓝 a)) (hg : Tendsto g x (𝓝 b)) : Tendsto (fun x => edist (f x) (g x)) x (𝓝 (edist a b)) := diff --git a/Mathlib/Topology/Instances/Int.lean b/Mathlib/Topology/Instances/Int.lean index 009674177e1f0..16ad9a3d16676 100644 --- a/Mathlib/Topology/Instances/Int.lean +++ b/Mathlib/Topology/Instances/Int.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro -/ import Mathlib.Data.Int.Interval -import Mathlib.Data.Int.SuccPred import Mathlib.Data.Int.ConditionallyCompleteOrder import Mathlib.Topology.Instances.Discrete import Mathlib.Topology.MetricSpace.Bounded @@ -73,7 +72,6 @@ theorem cobounded_eq : Bornology.cobounded ℤ = atBot ⊔ atTop := by simp_rw [← comap_dist_right_atTop (0 : ℤ), dist_eq', sub_zero, ← comap_abs_atTop, ← @Int.comap_cast_atTop ℝ, comap_comap]; rfl -@[deprecated (since := "2024-02-07")] alias cocompact_eq := cocompact_eq_atBot_atTop @[simp] theorem cofinite_eq : (cofinite : Filter ℤ) = atBot ⊔ atTop := by diff --git a/Mathlib/Topology/Instances/Irrational.lean b/Mathlib/Topology/Instances/Irrational.lean index f395be578fdff..8b82916e9b93b 100644 --- a/Mathlib/Topology/Instances/Irrational.lean +++ b/Mathlib/Topology/Instances/Irrational.lean @@ -36,7 +36,6 @@ open Filter Topology protected theorem IsGδ.setOf_irrational : IsGδ { x | Irrational x } := (countable_range _).isGδ_compl -@[deprecated (since := "2024-02-15")] alias isGδ_irrational := IsGδ.setOf_irrational theorem dense_irrational : Dense { x : ℝ | Irrational x } := by refine Real.isTopologicalBasis_Ioo_rat.dense_iff.2 ?_ diff --git a/Mathlib/Topology/Instances/Matrix.lean b/Mathlib/Topology/Instances/Matrix.lean index 6dbd675168fb4..13decc88c4c0a 100644 --- a/Mathlib/Topology/Instances/Matrix.lean +++ b/Mathlib/Topology/Instances/Matrix.lean @@ -263,7 +263,7 @@ variable [AddCommMonoid R] [TopologicalSpace R] theorem HasSum.matrix_transpose {f : X → Matrix m n R} {a : Matrix m n R} (hf : HasSum f a) : HasSum (fun x => (f x)ᵀ) aᵀ := - (hf.map (Matrix.transposeAddEquiv m n R) continuous_id.matrix_transpose : _) + (hf.map (Matrix.transposeAddEquiv m n R) continuous_id.matrix_transpose :) theorem Summable.matrix_transpose {f : X → Matrix m n R} (hf : Summable f) : Summable fun x => (f x)ᵀ := @@ -283,7 +283,7 @@ theorem Matrix.transpose_tsum [T2Space R] {f : X → Matrix m n R} : (∑' x, f theorem HasSum.matrix_conjTranspose [StarAddMonoid R] [ContinuousStar R] {f : X → Matrix m n R} {a : Matrix m n R} (hf : HasSum f a) : HasSum (fun x => (f x)ᴴ) aᴴ := - (hf.map (Matrix.conjTransposeAddEquiv m n R) continuous_id.matrix_conjTranspose : _) + (hf.map (Matrix.conjTransposeAddEquiv m n R) continuous_id.matrix_conjTranspose :) theorem Summable.matrix_conjTranspose [StarAddMonoid R] [ContinuousStar R] {f : X → Matrix m n R} (hf : Summable f) : Summable fun x => (f x)ᴴ := @@ -359,7 +359,7 @@ theorem Matrix.blockDiagonal_tsum [DecidableEq p] [T2Space R] {f : X → p → M theorem HasSum.matrix_blockDiag {f : X → Matrix (m × p) (n × p) R} {a : Matrix (m × p) (n × p) R} (hf : HasSum f a) : HasSum (fun x => blockDiag (f x)) (blockDiag a) := - (hf.map (blockDiagAddMonoidHom m n p R) <| Continuous.matrix_blockDiag continuous_id : _) + (hf.map (blockDiagAddMonoidHom m n p R) <| Continuous.matrix_blockDiag continuous_id :) theorem Summable.matrix_blockDiag {f : X → Matrix (m × p) (n × p) R} (hf : Summable f) : Summable fun x => blockDiag (f x) := diff --git a/Mathlib/Topology/Instances/Real.lean b/Mathlib/Topology/Instances/Real.lean index 522b40e629e4b..41f5e2e54a206 100644 --- a/Mathlib/Topology/Instances/Real.lean +++ b/Mathlib/Topology/Instances/Real.lean @@ -74,10 +74,7 @@ theorem Real.isTopologicalBasis_Ioo_rat : theorem Real.cobounded_eq : cobounded ℝ = atBot ⊔ atTop := by simp only [← comap_dist_right_atTop (0 : ℝ), Real.dist_eq, sub_zero, comap_abs_atTop] -@[deprecated (since := "2024-02-07")] alias Real.cocompact_eq := cocompact_eq_atBot_atTop -@[deprecated (since := "2024-02-07")] alias Real.atBot_le_cocompact := atBot_le_cocompact -@[deprecated (since := "2024-02-07")] alias Real.atTop_le_cocompact := atTop_le_cocompact /- TODO(Mario): Prove that these are uniform isomorphisms instead of uniform embeddings lemma uniform_embedding_add_rat {r : ℚ} : uniform_embedding (fun p : ℚ => p + r) := diff --git a/Mathlib/Topology/List.lean b/Mathlib/Topology/List.lean index c6bac03cf1d2e..be04e98105dbd 100644 --- a/Mathlib/Topology/List.lean +++ b/Mathlib/Topology/List.lean @@ -212,7 +212,7 @@ theorem continuous_insertIdx' {n : ℕ} {i : Fin (n + 1)} : theorem continuous_insertIdx {n : ℕ} {i : Fin (n + 1)} {f : β → α} {g : β → List.Vector α n} (hf : Continuous f) (hg : Continuous g) : Continuous fun b => Vector.insertIdx (f b) i (g b) := - continuous_insertIdx'.comp (hf.prod_mk hg : _) + continuous_insertIdx'.comp (hf.prod_mk hg) @[deprecated (since := "2024-10-21")] alias continuous_insertNth := continuous_insertIdx diff --git a/Mathlib/Topology/LocalAtTarget.lean b/Mathlib/Topology/LocalAtTarget.lean index af16705f2ce51..f1029b3df76df 100644 --- a/Mathlib/Topology/LocalAtTarget.lean +++ b/Mathlib/Topology/LocalAtTarget.lean @@ -195,7 +195,7 @@ To check that `f` is an embedding it suffices to check that `U i → Y` is an em theorem isEmbedding_of_iSup_eq_top_of_preimage_subset_range {X Y} [TopologicalSpace X] [TopologicalSpace Y] (f : X → Y) (h : Continuous f) {ι : Type*} - (U : ι → Opens Y) (hU : Set.range f ⊆ (iSup U : _)) + (U : ι → Opens Y) (hU : Set.range f ⊆ (iSup U :)) (V : ι → Type*) [∀ i, TopologicalSpace (V i)] (iV : ∀ i, V i → X) (hiV : ∀ i, Continuous (iV i)) (hV : ∀ i, f ⁻¹' U i ⊆ Set.range (iV i)) (hV' : ∀ i, IsEmbedding (f ∘ iV i)) : IsEmbedding f := by diff --git a/Mathlib/Topology/Maps/Basic.lean b/Mathlib/Topology/Maps/Basic.lean index 077466f86f877..ecd68d48bac01 100644 --- a/Mathlib/Topology/Maps/Basic.lean +++ b/Mathlib/Topology/Maps/Basic.lean @@ -487,7 +487,6 @@ theorem of_nonempty (h : ∀ s, IsClosed s → s.Nonempty → IsClosed (f '' s)) theorem isClosed_range (hf : IsClosedMap f) : IsClosed (range f) := @image_univ _ _ f ▸ hf _ isClosed_univ -@[deprecated (since := "2024-03-17")] alias closed_range := isClosed_range theorem isQuotientMap (hcl : IsClosedMap f) (hcont : Continuous f) (hsurj : Surjective f) : IsQuotientMap f := diff --git a/Mathlib/Topology/MetricSpace/Completion.lean b/Mathlib/Topology/MetricSpace/Completion.lean index cbca7d20888b3..e80a1ca356e08 100644 --- a/Mathlib/Topology/MetricSpace/Completion.lean +++ b/Mathlib/Topology/MetricSpace/Completion.lean @@ -45,7 +45,7 @@ protected theorem uniformContinuous_dist : /-- The new distance is continuous. -/ protected theorem continuous_dist [TopologicalSpace β] {f g : β → Completion α} (hf : Continuous f) (hg : Continuous g) : Continuous fun x ↦ dist (f x) (g x) := - Completion.uniformContinuous_dist.continuous.comp (hf.prod_mk hg : _) + Completion.uniformContinuous_dist.continuous.comp (hf.prod_mk hg :) /-- The new distance is an extension of the original distance. -/ @[simp] @@ -158,10 +158,6 @@ instance instMetricSpace : MetricSpace (Completion α) := toUniformSpace := inferInstance uniformity_dist := Completion.uniformity_dist } _ -@[deprecated eq_of_dist_eq_zero (since := "2024-03-10")] -protected theorem eq_of_dist_eq_zero (x y : Completion α) (h : dist x y = 0) : x = y := - eq_of_dist_eq_zero h - /-- The embedding of a metric space in its completion is an isometry. -/ theorem coe_isometry : Isometry ((↑) : α → Completion α) := Isometry.of_dist_eq Completion.dist_eq diff --git a/Mathlib/Topology/MetricSpace/PiNat.lean b/Mathlib/Topology/MetricSpace/PiNat.lean index b31d3608b2fed..db64438ea4873 100644 --- a/Mathlib/Topology/MetricSpace/PiNat.lean +++ b/Mathlib/Topology/MetricSpace/PiNat.lean @@ -681,7 +681,7 @@ theorem exists_nat_nat_continuous_surjective_of_completeSpace (α : Type*) [Metr let s : Set (ℕ → ℕ) := { x | (⋂ n : ℕ, closedBall (u (x n)) ((1 / 2) ^ n)).Nonempty } let g : s → α := fun x => x.2.some have A : ∀ (x : s) (n : ℕ), dist (g x) (u ((x : ℕ → ℕ) n)) ≤ (1 / 2) ^ n := fun x n => - (mem_iInter.1 x.2.some_mem n : _) + (mem_iInter.1 x.2.some_mem n :) have g_cont : Continuous g := by refine continuous_iff_continuousAt.2 fun y => ?_ refine continuousAt_of_locally_lipschitz zero_lt_one 4 fun x hxy => ?_ diff --git a/Mathlib/Topology/MetricSpace/Polish.lean b/Mathlib/Topology/MetricSpace/Polish.lean index 70ae87a47d620..543d4da8c0d76 100644 --- a/Mathlib/Topology/MetricSpace/Polish.lean +++ b/Mathlib/Topology/MetricSpace/Polish.lean @@ -102,9 +102,6 @@ instance (priority := 100) instMetrizableSpace (α : Type*) [TopologicalSpace α letI := upgradePolishSpace α infer_instance -@[deprecated "No deprecation message was provided." (since := "2024-02-23")] -theorem t2Space (α : Type*) [TopologicalSpace α] [PolishSpace α] : T2Space α := inferInstance - /-- A countable product of Polish spaces is Polish. -/ instance pi_countable {ι : Type*} [Countable ι] {E : ι → Type*} [∀ i, TopologicalSpace (E i)] [∀ i, PolishSpace (E i)] : PolishSpace (∀ i, E i) := by diff --git a/Mathlib/Topology/MetricSpace/Pseudo/Constructions.lean b/Mathlib/Topology/MetricSpace/Pseudo/Constructions.lean index aba39bbafa59c..844690eff5d58 100644 --- a/Mathlib/Topology/MetricSpace/Pseudo/Constructions.lean +++ b/Mathlib/Topology/MetricSpace/Pseudo/Constructions.lean @@ -207,7 +207,7 @@ lemma continuous_dist : Continuous fun p : α × α ↦ dist p.1 p.2 := uniformC @[continuity, fun_prop] protected lemma Continuous.dist [TopologicalSpace β] {f g : β → α} (hf : Continuous f) (hg : Continuous g) : Continuous fun b => dist (f b) (g b) := - continuous_dist.comp (hf.prod_mk hg : _) + continuous_dist.comp (hf.prod_mk hg :) protected lemma Filter.Tendsto.dist {f g : β → α} {x : Filter β} {a b : α} (hf : Tendsto f x (𝓝 a)) (hg : Tendsto g x (𝓝 b)) : @@ -233,7 +233,7 @@ lemma continuous_nndist : Continuous fun p : α × α => nndist p.1 p.2 := @[fun_prop] protected lemma Continuous.nndist [TopologicalSpace β] {f g : β → α} (hf : Continuous f) (hg : Continuous g) : Continuous fun b => nndist (f b) (g b) := - continuous_nndist.comp (hf.prod_mk hg : _) + continuous_nndist.comp (hf.prod_mk hg :) protected lemma Filter.Tendsto.nndist {f g : β → α} {x : Filter β} {a b : α} (hf : Tendsto f x (𝓝 a)) (hg : Tendsto g x (𝓝 b)) : diff --git a/Mathlib/Topology/MetricSpace/Sequences.lean b/Mathlib/Topology/MetricSpace/Sequences.lean index ed02f21cdebd7..5548bfac6db34 100644 --- a/Mathlib/Topology/MetricSpace/Sequences.lean +++ b/Mathlib/Topology/MetricSpace/Sequences.lean @@ -17,12 +17,6 @@ open scoped Topology variable {X : Type*} [PseudoMetricSpace X] -@[deprecated lebesgue_number_lemma_of_metric (since := "2024-02-24")] -nonrec theorem SeqCompact.lebesgue_number_lemma_of_metric {ι : Sort*} {c : ι → Set X} {s : Set X} - (hs : IsSeqCompact s) (hc₁ : ∀ i, IsOpen (c i)) (hc₂ : s ⊆ ⋃ i, c i) : - ∃ δ > 0, ∀ a ∈ s, ∃ i, ball a δ ⊆ c i := - lebesgue_number_lemma_of_metric hs.isCompact hc₁ hc₂ - variable [ProperSpace X] {s : Set X} /-- A version of **Bolzano-Weierstrass**: in a proper metric space (eg. $ℝ^n$), diff --git a/Mathlib/Topology/MetricSpace/Thickening.lean b/Mathlib/Topology/MetricSpace/Thickening.lean index 8ef9ed0bb9255..67b4aa3b257bc 100644 --- a/Mathlib/Topology/MetricSpace/Thickening.lean +++ b/Mathlib/Topology/MetricSpace/Thickening.lean @@ -99,6 +99,7 @@ theorem frontier_thickening_subset (E : Set α) {δ : ℝ} : frontier (thickening δ E) ⊆ { x : α | infEdist x E = ENNReal.ofReal δ } := frontier_lt_subset_eq continuous_infEdist continuous_const +open scoped Function in -- required for scoped `on` notation theorem frontier_thickening_disjoint (A : Set α) : Pairwise (Disjoint on fun r : ℝ => frontier (thickening r A)) := by refine (pairwise_disjoint_on _).2 fun r₁ r₂ hr => ?_ @@ -622,6 +623,7 @@ theorem cthickening_cthickening_subset (hε : 0 ≤ ε) (hδ : 0 ≤ δ) (s : Se simp_rw [mem_cthickening_iff, ENNReal.ofReal_add hε hδ] exact fun hx => infEdist_le_infEdist_cthickening_add.trans (add_le_add_right hx _) +open scoped Function in -- required for scoped `on` notation theorem frontier_cthickening_disjoint (A : Set α) : Pairwise (Disjoint on fun r : ℝ≥0 => frontier (cthickening r A)) := fun r₁ r₂ hr => ((disjoint_singleton.2 <| by simpa).preimage _).mono (frontier_cthickening_subset _) diff --git a/Mathlib/Topology/Order.lean b/Mathlib/Topology/Order.lean index 0be6eb00b68ce..2bd638f106979 100644 --- a/Mathlib/Topology/Order.lean +++ b/Mathlib/Topology/Order.lean @@ -286,9 +286,6 @@ theorem le_of_nhds_le_nhds (h : ∀ x, @nhds α t₁ x ≤ @nhds α t₂ x) : t rw [@isOpen_iff_mem_nhds _ _ t₁, @isOpen_iff_mem_nhds α _ t₂] exact fun hs a ha => h _ (hs _ ha) -@[deprecated (since := "2024-03-01")] -alias eq_of_nhds_eq_nhds := TopologicalSpace.ext_nhds - theorem eq_bot_of_singletons_open {t : TopologicalSpace α} (h : ∀ x, IsOpen[t] {x}) : t = ⊥ := bot_unique fun s _ => biUnion_of_singleton s ▸ isOpen_biUnion fun x _ => h x @@ -572,19 +569,11 @@ theorem nhds_nhdsAdjoint_same (a : α) (f : Filter α) : exact IsOpen.mem_nhds (fun _ ↦ htf) hat · exact sup_le (pure_le_nhds _) ((gc_nhds a).le_u_l f) -@[deprecated (since := "2024-02-10")] -alias nhdsAdjoint_nhds := nhds_nhdsAdjoint_same - theorem nhds_nhdsAdjoint_of_ne {a b : α} (f : Filter α) (h : b ≠ a) : @nhds α (nhdsAdjoint a f) b = pure b := let _ := nhdsAdjoint a f (isOpen_singleton_iff_nhds_eq_pure _).1 <| isOpen_singleton_nhdsAdjoint f h -@[deprecated nhds_nhdsAdjoint_of_ne (since := "2024-02-10")] -theorem nhdsAdjoint_nhds_of_ne (a : α) (f : Filter α) {b : α} (h : b ≠ a) : - @nhds α (nhdsAdjoint a f) b = pure b := - nhds_nhdsAdjoint_of_ne f h - theorem nhds_nhdsAdjoint [DecidableEq α] (a : α) (f : Filter α) : @nhds α (nhdsAdjoint a f) = update pure a (pure a ⊔ f) := eq_update_iff.2 ⟨nhds_nhdsAdjoint_same .., fun _ ↦ nhds_nhdsAdjoint_of_ne _⟩ diff --git a/Mathlib/Topology/Order/Basic.lean b/Mathlib/Topology/Order/Basic.lean index a0884956957dd..80dc91a3e3e3e 100644 --- a/Mathlib/Topology/Order/Basic.lean +++ b/Mathlib/Topology/Order/Basic.lean @@ -489,16 +489,6 @@ theorem Dense.topology_eq_generateFrom [OrderTopology α] [DenselyOrdered α] {s let _ := generateFrom (Ioi '' s ∪ Iio '' s) exact isOpen_iUnion fun x ↦ isOpen_iUnion fun h ↦ .basic _ <| .inr <| mem_image_of_mem _ h.1 -@[deprecated OrderBot.atBot_eq (since := "2024-02-14")] -theorem atBot_le_nhds_bot [OrderBot α] : (atBot : Filter α) ≤ 𝓝 ⊥ := by - rw [OrderBot.atBot_eq] - apply pure_le_nhds - -set_option linter.deprecated false in -@[deprecated OrderTop.atTop_eq (since := "2024-02-14")] -theorem atTop_le_nhds_top [OrderTop α] : (atTop : Filter α) ≤ 𝓝 ⊤ := - @atBot_le_nhds_bot αᵒᵈ _ _ _ - variable (α) /-- Let `α` be a densely ordered linear order with order topology. If `α` is a separable space, then diff --git a/Mathlib/Topology/Order/DenselyOrdered.lean b/Mathlib/Topology/Order/DenselyOrdered.lean index 0c981229bfc16..957c94d57c1e0 100644 --- a/Mathlib/Topology/Order/DenselyOrdered.lean +++ b/Mathlib/Topology/Order/DenselyOrdered.lean @@ -259,9 +259,9 @@ theorem map_coe_atTop_of_Ioo_subset (hb : s ⊆ Iio b) (hs : ∀ a' < b, ∃ a < theorem map_coe_atBot_of_Ioo_subset (ha : s ⊆ Ioi a) (hs : ∀ b' > a, ∃ b > a, Ioo a b ⊆ s) : map ((↑) : s → α) atBot = 𝓝[>] a := by - -- the elaborator gets stuck without `(... : _)` + -- the elaborator gets stuck without `(... :)` refine (map_coe_atTop_of_Ioo_subset (show ofDual ⁻¹' s ⊆ Iio (toDual a) from ha) - fun b' hb' => ?_ : _) + fun b' hb' => ?_ :) simpa only [OrderDual.exists, dual_Ioo] using hs b' hb' /-- The `atTop` filter for an open interval `Ioo a b` comes from the left-neighbourhoods filter at diff --git a/Mathlib/Topology/Order/IsLocallyClosed.lean b/Mathlib/Topology/Order/IsLocallyClosed.lean new file mode 100644 index 0000000000000..675c11d32edb6 --- /dev/null +++ b/Mathlib/Topology/Order/IsLocallyClosed.lean @@ -0,0 +1,49 @@ +/- +Copyright (c) 2024 Xavier Roblot. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Xavier Roblot +-/ +import Mathlib.Topology.Order.OrderClosed +import Mathlib.Topology.LocallyClosed + +/-! +# Intervals are locally closed + +We prove that the intervals on a topological ordered space are locally closed. +-/ + +variable {X : Type*} [TopologicalSpace X] {a b : X} + +theorem isLocallyClosed_Icc [Preorder X] [OrderClosedTopology X] : + IsLocallyClosed (Set.Icc a b) := + isClosed_Icc.isLocallyClosed + +theorem isLocallyClosed_Ioo [LinearOrder X] [OrderClosedTopology X] : + IsLocallyClosed (Set.Ioo a b) := + isOpen_Ioo.isLocallyClosed + +theorem isLocallyClosed_Ici [Preorder X] [ClosedIciTopology X] : + IsLocallyClosed (Set.Ici a) := + isClosed_Ici.isLocallyClosed + +theorem isLocallyClosed_Iic [Preorder X] [ClosedIicTopology X] : + IsLocallyClosed (Set.Iic a) := + isClosed_Iic.isLocallyClosed + +theorem isLocallyClosed_Ioi [LinearOrder X] [ClosedIicTopology X] : + IsLocallyClosed (Set.Ioi a) := + isOpen_Ioi.isLocallyClosed + +theorem isLocallyClosed_Iio [LinearOrder X] [ClosedIciTopology X] : + IsLocallyClosed (Set.Iio a) := + isOpen_Iio.isLocallyClosed + +theorem isLocallyClosed_Ioc [LinearOrder X] [ClosedIicTopology X] : + IsLocallyClosed (Set.Ioc a b) := by + rw [← Set.Iic_inter_Ioi] + exact isLocallyClosed_Iic.inter isLocallyClosed_Ioi + +theorem isLocallyClosed_Ico [LinearOrder X] [ClosedIciTopology X] : + IsLocallyClosed (Set.Ico a b) := by + rw [← Set.Iio_inter_Ici] + exact isLocallyClosed_Iio.inter isLocallyClosed_Ici diff --git a/Mathlib/Topology/Order/Lattice.lean b/Mathlib/Topology/Order/Lattice.lean index 1ef8035f3ab65..60a021ba8a061 100644 --- a/Mathlib/Topology/Order/Lattice.lean +++ b/Mathlib/Topology/Order/Lattice.lean @@ -77,7 +77,7 @@ theorem continuous_inf [Min L] [ContinuousInf L] : Continuous fun p : L × L => @[continuity, fun_prop] theorem Continuous.inf [Min L] [ContinuousInf L] {f g : X → L} (hf : Continuous f) (hg : Continuous g) : Continuous fun x => f x ⊓ g x := - continuous_inf.comp (hf.prod_mk hg : _) + continuous_inf.comp (hf.prod_mk hg :) @[continuity] theorem continuous_sup [Max L] [ContinuousSup L] : Continuous fun p : L × L => p.1 ⊔ p.2 := @@ -86,7 +86,7 @@ theorem continuous_sup [Max L] [ContinuousSup L] : Continuous fun p : L × L => @[continuity, fun_prop] theorem Continuous.sup [Max L] [ContinuousSup L] {f g : X → L} (hf : Continuous f) (hg : Continuous g) : Continuous fun x => f x ⊔ g x := - continuous_sup.comp (hf.prod_mk hg : _) + continuous_sup.comp (hf.prod_mk hg :) namespace Filter.Tendsto diff --git a/Mathlib/Topology/Order/LowerUpperTopology.lean b/Mathlib/Topology/Order/LowerUpperTopology.lean index 34ccdbc54eeb5..2f58eb8afa600 100644 --- a/Mathlib/Topology/Order/LowerUpperTopology.lean +++ b/Mathlib/Topology/Order/LowerUpperTopology.lean @@ -102,16 +102,32 @@ protected def rec {β : WithLower α → Sort*} (h : ∀ a, β (toLower a)) : instance [Nonempty α] : Nonempty (WithLower α) := ‹Nonempty α› instance [Inhabited α] : Inhabited (WithLower α) := ‹Inhabited α› +section Preorder + variable [Preorder α] {s : Set α} instance : Preorder (WithLower α) := ‹Preorder α› -instance : TopologicalSpace (WithLower α) := lower α +instance : TopologicalSpace (WithLower α) := lower (WithLower α) + +@[simp] lemma toLower_le_toLower {x y : α} : toLower x ≤ toLower y ↔ x ≤ y := .rfl +@[simp] lemma toLower_lt_toLower {x y : α} : toLower x < toLower y ↔ x < y := .rfl +@[simp] lemma ofLower_le_ofLower {x y : WithLower α} : ofLower x ≤ ofLower y ↔ x ≤ y := .rfl +@[simp] lemma ofLower_lt_ofLower {x y : WithLower α} : ofLower x < ofLower y ↔ x < y := .rfl -lemma isOpen_preimage_ofLower : IsOpen (ofLower ⁻¹' s) ↔ (lower α).IsOpen s := Iff.rfl +lemma isOpen_preimage_ofLower : IsOpen (ofLower ⁻¹' s) ↔ IsOpen[lower α] s := Iff.rfl -lemma isOpen_def (T : Set (WithLower α)) : IsOpen T ↔ (lower α).IsOpen (WithLower.toLower ⁻¹' T) := +lemma isOpen_def (T : Set (WithLower α)) : IsOpen T ↔ IsOpen[lower α] (WithLower.toLower ⁻¹' T) := Iff.rfl +theorem continuous_toLower [TopologicalSpace α] [ClosedIciTopology α] : + Continuous (toLower : α → WithLower α) := + continuous_generateFrom_iff.mpr <| by rintro _ ⟨a, rfl⟩; exact isClosed_Ici.isOpen_compl + +end Preorder + +instance [PartialOrder α] : PartialOrder (WithLower α) := ‹PartialOrder α› +instance [LinearOrder α] : LinearOrder (WithLower α) := ‹LinearOrder α› + end WithLower /-- Type synonym for a preorder equipped with the upper topology. -/ @@ -141,15 +157,31 @@ protected def rec {β : WithUpper α → Sort*} (h : ∀ a, β (toUpper a)) : instance [Nonempty α] : Nonempty (WithUpper α) := ‹Nonempty α› instance [Inhabited α] : Inhabited (WithUpper α) := ‹Inhabited α› +section Preorder + variable [Preorder α] {s : Set α} instance : Preorder (WithUpper α) := ‹Preorder α› -instance : TopologicalSpace (WithUpper α) := upper α +instance : TopologicalSpace (WithUpper α) := upper (WithUpper α) + +@[simp] lemma toUpper_le_toUpper {x y : α} : toUpper x ≤ toUpper y ↔ x ≤ y := .rfl +@[simp] lemma toUpper_lt_toUpper {x y : α} : toUpper x < toUpper y ↔ x < y := .rfl +@[simp] lemma ofUpper_le_ofUpper {x y : WithUpper α} : ofUpper x ≤ ofUpper y ↔ x ≤ y := .rfl +@[simp] lemma ofUpper_lt_ofUpper {x y : WithUpper α} : ofUpper x < ofUpper y ↔ x < y := .rfl lemma isOpen_preimage_ofUpper : IsOpen (ofUpper ⁻¹' s) ↔ (upper α).IsOpen s := Iff.rfl lemma isOpen_def {s : Set (WithUpper α)} : IsOpen s ↔ (upper α).IsOpen (toUpper ⁻¹' s) := Iff.rfl +theorem continuous_toUpper [TopologicalSpace α] [ClosedIicTopology α] : + Continuous (toUpper : α → WithUpper α) := + continuous_generateFrom_iff.mpr <| by rintro _ ⟨a, rfl⟩; exact isClosed_Iic.isOpen_compl + +end Preorder + +instance [PartialOrder α] : PartialOrder (WithUpper α) := ‹PartialOrder α› +instance [LinearOrder α] : LinearOrder (WithUpper α) := ‹LinearOrder α› + end WithUpper /-- @@ -236,6 +268,11 @@ theorem isLowerSet_of_isOpen (h : IsOpen s) : IsLowerSet s := by theorem isUpperSet_of_isClosed (h : IsClosed s) : IsUpperSet s := isLowerSet_compl.1 <| isLowerSet_of_isOpen h.isOpen_compl +theorem tendsto_nhds_iff_not_le {β : Type*} {f : β → α} {l : Filter β} {x : α} : + Filter.Tendsto f l (𝓝 x) ↔ ∀ y, ¬y ≤ x → ∀ᶠ z in l, ¬y ≤ f z := by + simp [topology_eq_lowerTopology, tendsto_nhds_generateFrom_iff, Filter.Eventually, Ici, + compl_setOf] + /-- The closure of a singleton `{a}` in the lower topology is the left-closed right-infinite interval [a, ∞). @@ -292,6 +329,10 @@ lemma isTopologicalBasis_insert_univ_subbasis : use b ⊓ c rw [compl_Ici, compl_Ici, compl_Ici, Iio_inter_Iio]) +theorem tendsto_nhds_iff_lt {β : Type*} {f : β → α} {l : Filter β} {x : α} : + Filter.Tendsto f l (𝓝 x) ↔ ∀ y, x < y → ∀ᶠ z in l, f z < y := by + simp only [tendsto_nhds_iff_not_le, not_le] + end LinearOrder section CompleteLinearOrder @@ -382,6 +423,10 @@ theorem isUpperSet_of_isOpen (h : IsOpen s) : IsUpperSet s := theorem isLowerSet_of_isClosed (h : IsClosed s) : IsLowerSet s := isUpperSet_compl.1 <| isUpperSet_of_isOpen h.isOpen_compl +theorem tendsto_nhds_iff_not_le {β : Type*} {f : β → α} {l : Filter β} {x : α} : + Filter.Tendsto f l (𝓝 x) ↔ ∀ y, ¬x ≤ y → ∀ᶠ z in l, ¬f z ≤ y := + IsLower.tendsto_nhds_iff_not_le (α := αᵒᵈ) + /-- The closure of a singleton `{a}` in the upper topology is the left-infinite right-closed interval (-∞,a]. @@ -421,6 +466,10 @@ lemma isTopologicalBasis_insert_univ_subbasis : IsTopologicalBasis (insert univ {s : Set α | ∃ a, (Iic a)ᶜ = s}) := IsLower.isTopologicalBasis_insert_univ_subbasis (α := αᵒᵈ) +theorem tendsto_nhds_iff_lt {β : Type*} {f : β → α} {l : Filter β} {x : α} : + Filter.Tendsto f l (𝓝 x) ↔ ∀ y < x, ∀ᶠ z in l, y < f z := + IsLower.tendsto_nhds_iff_lt (α := αᵒᵈ) + end LinearOrder section CompleteLinearOrder diff --git a/Mathlib/Topology/Order/OrderClosed.lean b/Mathlib/Topology/Order/OrderClosed.lean index 94bdce684cd4a..846125f18cf39 100644 --- a/Mathlib/Topology/Order/OrderClosed.lean +++ b/Mathlib/Topology/Order/OrderClosed.lean @@ -108,10 +108,6 @@ variable [TopologicalSpace α] [Preorder α] [ClosedIicTopology α] {f : β → theorem isClosed_Iic : IsClosed (Iic a) := ClosedIicTopology.isClosed_Iic a -@[deprecated isClosed_Iic (since := "2024-02-15")] -lemma ClosedIicTopology.isClosed_le' (a : α) : IsClosed {x | x ≤ a} := isClosed_Iic a -export ClosedIicTopology (isClosed_le') - instance : ClosedIciTopology αᵒᵈ where isClosed_Ici _ := isClosed_Iic (α := α) @@ -415,10 +411,6 @@ variable [TopologicalSpace α] [Preorder α] [ClosedIciTopology α] {f : β → theorem isClosed_Ici {a : α} : IsClosed (Ici a) := ClosedIciTopology.isClosed_Ici a -@[deprecated isClosed_Ici (since := "2024-02-15")] -lemma ClosedIciTopology.isClosed_ge' (a : α) : IsClosed {x | a ≤ x} := isClosed_Ici a -export ClosedIciTopology (isClosed_ge') - instance : ClosedIicTopology αᵒᵈ where isClosed_Iic _ := isClosed_Ici (α := α) diff --git a/Mathlib/Topology/PartialHomeomorph.lean b/Mathlib/Topology/PartialHomeomorph.lean index 0363f5b15a445..20435dc149290 100644 --- a/Mathlib/Topology/PartialHomeomorph.lean +++ b/Mathlib/Topology/PartialHomeomorph.lean @@ -1018,7 +1018,7 @@ theorem continuousWithinAt_iff_continuousWithinAt_comp_left {f : Z → X} {s : S neighborhood of the initial point is sent to the source of the partial homeomorphism -/ theorem continuousAt_iff_continuousAt_comp_left {f : Z → X} {x : Z} (h : f ⁻¹' e.source ∈ 𝓝 x) : ContinuousAt f x ↔ ContinuousAt (e ∘ f) x := by - have hx : f x ∈ e.source := (mem_of_mem_nhds h : _) + have hx : f x ∈ e.source := (mem_of_mem_nhds h :) have h' : f ⁻¹' e.source ∈ 𝓝[univ] x := by rwa [nhdsWithin_univ] rw [← continuousWithinAt_univ, ← continuousWithinAt_univ, e.continuousWithinAt_iff_continuousWithinAt_comp_left hx h'] diff --git a/Mathlib/Topology/Semicontinuous.lean b/Mathlib/Topology/Semicontinuous.lean index ae630cd687ba7..b569bcd00df4f 100644 --- a/Mathlib/Topology/Semicontinuous.lean +++ b/Mathlib/Topology/Semicontinuous.lean @@ -326,14 +326,8 @@ theorem lowerSemicontinuous_iff_isClosed_epigraph {f : α → γ} : · rw [lowerSemicontinuous_iff_isClosed_preimage] exact fun hf y ↦ hf.preimage (Continuous.Prod.mk_left y) -@[deprecated (since := "2024-03-02")] -alias lowerSemicontinuous_iff_IsClosed_epigraph := lowerSemicontinuous_iff_isClosed_epigraph - alias ⟨LowerSemicontinuous.isClosed_epigraph, _⟩ := lowerSemicontinuous_iff_isClosed_epigraph -@[deprecated (since := "2024-03-02")] -alias LowerSemicontinuous.IsClosed_epigraph := LowerSemicontinuous.isClosed_epigraph - end /-! ### Composition -/ diff --git a/Mathlib/Topology/Separation/GDelta.lean b/Mathlib/Topology/Separation/GDelta.lean index abc821f17ade6..57bc39f4a2d32 100644 --- a/Mathlib/Topology/Separation/GDelta.lean +++ b/Mathlib/Topology/Separation/GDelta.lean @@ -35,7 +35,6 @@ section Separation theorem IsGδ.compl_singleton (x : X) [T1Space X] : IsGδ ({x}ᶜ : Set X) := isOpen_compl_singleton.isGδ -@[deprecated (since := "2024-02-15")] alias isGδ_compl_singleton := IsGδ.compl_singleton theorem Set.Countable.isGδ_compl {s : Set X} [T1Space X] (hs : s.Countable) : IsGδ sᶜ := by rw [← biUnion_of_singleton s, compl_iUnion₂] @@ -56,7 +55,6 @@ protected theorem IsGδ.singleton [FirstCountableTopology X] [T1Space X] (x : X) rw [← biInter_basis_nhds h_basis.toHasBasis] exact .biInter (to_countable _) fun n _ => (hU n).2.isGδ -@[deprecated (since := "2024-02-15")] alias isGδ_singleton := IsGδ.singleton theorem Set.Finite.isGδ [FirstCountableTopology X] {s : Set X} [T1Space X] (hs : s.Finite) : IsGδ s := diff --git a/Mathlib/Topology/Separation/Hausdorff.lean b/Mathlib/Topology/Separation/Hausdorff.lean index 17cc8d6659b95..957b2b4b1cfd1 100644 --- a/Mathlib/Topology/Separation/Hausdorff.lean +++ b/Mathlib/Topology/Separation/Hausdorff.lean @@ -528,9 +528,6 @@ theorem Function.LeftInverse.isClosed_range [T2Space X] {f : X → Y} {g : Y → h.rightInvOn_range.eqOn.closure (hg.comp hf) continuous_id isClosed_of_closure_subset fun x hx => ⟨f x, this hx⟩ -@[deprecated (since := "2024-03-17")] -alias Function.LeftInverse.closed_range := Function.LeftInverse.isClosed_range - theorem Function.LeftInverse.isClosedEmbedding [T2Space X] {f : X → Y} {g : Y → X} (h : Function.LeftInverse f g) (hf : Continuous f) (hg : Continuous g) : IsClosedEmbedding g := ⟨.of_leftInverse h hf hg, h.isClosed_range hf hg⟩ diff --git a/Mathlib/Topology/Separation/Regular.lean b/Mathlib/Topology/Separation/Regular.lean index f886e9d275213..3be8602edb54b 100644 --- a/Mathlib/Topology/Separation/Regular.lean +++ b/Mathlib/Topology/Separation/Regular.lean @@ -113,24 +113,15 @@ theorem RegularSpace.of_lift'_closure_le (h : ∀ x : X, (𝓝 x).lift' closure theorem RegularSpace.of_lift'_closure (h : ∀ x : X, (𝓝 x).lift' closure = 𝓝 x) : RegularSpace X := Iff.mpr ((regularSpace_TFAE X).out 0 5) h -@[deprecated (since := "2024-02-28")] -alias RegularSpace.ofLift'_closure := RegularSpace.of_lift'_closure - theorem RegularSpace.of_hasBasis {ι : X → Sort*} {p : ∀ a, ι a → Prop} {s : ∀ a, ι a → Set X} (h₁ : ∀ a, (𝓝 a).HasBasis (p a) (s a)) (h₂ : ∀ a i, p a i → IsClosed (s a i)) : RegularSpace X := .of_lift'_closure fun a => (h₁ a).lift'_closure_eq_self (h₂ a) -@[deprecated (since := "2024-02-28")] -alias RegularSpace.ofBasis := RegularSpace.of_hasBasis - theorem RegularSpace.of_exists_mem_nhds_isClosed_subset (h : ∀ (x : X), ∀ s ∈ 𝓝 x, ∃ t ∈ 𝓝 x, IsClosed t ∧ t ⊆ s) : RegularSpace X := Iff.mpr ((regularSpace_TFAE X).out 0 3) h -@[deprecated (since := "2024-02-28")] -alias RegularSpace.ofExistsMemNhdsIsClosedSubset := RegularSpace.of_exists_mem_nhds_isClosed_subset - /-- A weakly locally compact R₁ space is regular. -/ instance (priority := 100) [WeaklyLocallyCompactSpace X] [R1Space X] : RegularSpace X := .of_hasBasis isCompact_isClosed_basis_nhds fun _ _ ⟨_, _, h⟩ ↦ h diff --git a/Mathlib/Topology/Sheaves/SheafCondition/OpensLeCover.lean b/Mathlib/Topology/Sheaves/SheafCondition/OpensLeCover.lean index 318e4ad961054..d316325a6727d 100644 --- a/Mathlib/Topology/Sheaves/SheafCondition/OpensLeCover.lean +++ b/Mathlib/Topology/Sheaves/SheafCondition/OpensLeCover.lean @@ -154,7 +154,7 @@ def generateEquivalenceOpensLe (hY : Y = iSup U) : inverse := generateEquivalenceOpensLe_inverse' _ hY unitIso := eqToIso <| CategoryTheory.Functor.ext (by rintro ⟨⟨_, _⟩, _⟩; dsimp; congr) - (by intros; refine Over.OverMorphism.ext ?_; aesop_cat) + (by intros; refine Over.OverMorphism.ext ?_; simp) counitIso := eqToIso <| CategoryTheory.Functor.hext (by intro; refine FullSubcategory.ext ?_; rfl) (by intros; rfl) diff --git a/Mathlib/Topology/Sheaves/SheafCondition/PairwiseIntersections.lean b/Mathlib/Topology/Sheaves/SheafCondition/PairwiseIntersections.lean index 4ae09d65bbf39..807fbd616d09b 100644 --- a/Mathlib/Topology/Sheaves/SheafCondition/PairwiseIntersections.lean +++ b/Mathlib/Topology/Sheaves/SheafCondition/PairwiseIntersections.lean @@ -217,7 +217,7 @@ after appropriate whiskering and postcomposition. -/ def pairwiseCoconeIso : (Pairwise.cocone U).op ≅ - (Cones.postcomposeEquivalence (NatIso.op (pairwiseDiagramIso U : _) : _)).functor.obj + (Cones.postcomposeEquivalence (NatIso.op (pairwiseDiagramIso U :) :)).functor.obj ((opensLeCoverCocone U).op.whisker (pairwiseToOpensLeCover U).op) := Cones.ext (Iso.refl _) (by aesop_cat) @@ -254,7 +254,7 @@ theorem isSheafOpensLeCover_iff_isSheafPairwiseIntersections : ((opensLeCoverCocone U).op.whisker (pairwiseToOpensLeCover U).op))) := (IsLimit.equivIsoLimit (Functor.mapConePostcomposeEquivalenceFunctor _).symm) _ ≃ IsLimit (F.mapCone (Pairwise.cocone U).op) := - IsLimit.equivIsoLimit ((Cones.functoriality _ _).mapIso (pairwiseCoconeIso U : _).symm) + IsLimit.equivIsoLimit ((Cones.functoriality _ _).mapIso (pairwiseCoconeIso U :).symm) /-- The sheaf condition in terms of an equalizer diagram is equivalent to the reformulation in terms of a limit diagram over `U i` and `U i ⊓ U j`. diff --git a/Mathlib/Topology/Sheaves/Stalks.lean b/Mathlib/Topology/Sheaves/Stalks.lean index 28c58d23c6a6d..3b24d1d5a253d 100644 --- a/Mathlib/Topology/Sheaves/Stalks.lean +++ b/Mathlib/Topology/Sheaves/Stalks.lean @@ -335,12 +335,12 @@ noncomputable def stalkSpecializes (F : X.Presheaf C) {x y : X} (h : x ⤳ y) : refine colimit.desc _ ⟨_, fun U => ?_, ?_⟩ · exact colimit.ι ((OpenNhds.inclusion x).op ⋙ F) - (op ⟨(unop U).1, (specializes_iff_forall_open.mp h _ (unop U).1.2 (unop U).2 : _)⟩) + (op ⟨(unop U).1, (specializes_iff_forall_open.mp h _ (unop U).1.2 (unop U).2 :)⟩) · intro U V i dsimp rw [Category.comp_id] - let U' : OpenNhds x := ⟨_, (specializes_iff_forall_open.mp h _ (unop U).1.2 (unop U).2 : _)⟩ - let V' : OpenNhds x := ⟨_, (specializes_iff_forall_open.mp h _ (unop V).1.2 (unop V).2 : _)⟩ + let U' : OpenNhds x := ⟨_, (specializes_iff_forall_open.mp h _ (unop U).1.2 (unop U).2 :)⟩ + let V' : OpenNhds x := ⟨_, (specializes_iff_forall_open.mp h _ (unop V).1.2 (unop V).2 :)⟩ exact colimit.w ((OpenNhds.inclusion x).op ⋙ F) (show V' ⟶ U' from i.unop).op @[reassoc (attr := simp), elementwise nosimp] diff --git a/Mathlib/Topology/UniformSpace/AbstractCompletion.lean b/Mathlib/Topology/UniformSpace/AbstractCompletion.lean index d5458eff1e284..67a5460ca8634 100644 --- a/Mathlib/Topology/UniformSpace/AbstractCompletion.lean +++ b/Mathlib/Topology/UniformSpace/AbstractCompletion.lean @@ -374,7 +374,7 @@ theorem uniformContinuous_map₂ (f : α → β → γ) : UniformContinuous₂ ( theorem continuous_map₂ {δ} [TopologicalSpace δ] {f : α → β → γ} {a : δ → hatα} {b : δ → hatβ} (ha : Continuous a) (hb : Continuous b) : Continuous fun d : δ => pkg.map₂ pkg' pkg'' f (a d) (b d) := - ((pkg.uniformContinuous_map₂ pkg' pkg'' f).continuous.comp (Continuous.prod_mk ha hb) : _) + ((pkg.uniformContinuous_map₂ pkg' pkg'' f).continuous.comp (Continuous.prod_mk ha hb) :) theorem map₂_coe_coe (a : α) (b : β) (f : α → β → γ) (hf : UniformContinuous₂ f) : pkg.map₂ pkg' pkg'' f (ι a) (ι' b) = ι'' (f a b) := diff --git a/Mathlib/Topology/UniformSpace/Basic.lean b/Mathlib/Topology/UniformSpace/Basic.lean index 0057d6c5ff78f..f5abe4496c48a 100644 --- a/Mathlib/Topology/UniformSpace/Basic.lean +++ b/Mathlib/Topology/UniformSpace/Basic.lean @@ -352,20 +352,6 @@ theorem UniformSpace.toCore_toTopologicalSpace (u : UniformSpace α) : TopologicalSpace.ext_nhds fun a ↦ by rw [u.nhds_eq_comap_uniformity, u.toCore.nhds_toTopologicalSpace] -/-- Build a `UniformSpace` from a `UniformSpace.Core` and a compatible topology. -Use `UniformSpace.mk` instead to avoid proving -the unnecessary assumption `UniformSpace.Core.refl`. - -The main constructor used to use a different compatibility assumption. -This definition was created as a step towards porting to a new definition. -Now the main definition is ported, -so this constructor will be removed in a few months. -/ -@[deprecated UniformSpace.mk (since := "2024-03-20")] -def UniformSpace.ofNhdsEqComap (u : UniformSpace.Core α) (_t : TopologicalSpace α) - (h : ∀ x, 𝓝 x = u.uniformity.comap (Prod.mk x)) : UniformSpace α where - __ := u - nhds_eq_comap_uniformity := h - @[ext (iff := false)] protected theorem UniformSpace.ext {u₁ u₂ : UniformSpace α} (h : 𝓤[u₁] = 𝓤[u₂]) : u₁ = u₂ := by have : u₁.toTopologicalSpace = u₂.toTopologicalSpace := TopologicalSpace.ext_nhds fun x ↦ by diff --git a/Mathlib/Topology/UniformSpace/Cauchy.lean b/Mathlib/Topology/UniformSpace/Cauchy.lean index 203ebe020f1ac..18f55b5f95253 100644 --- a/Mathlib/Topology/UniformSpace/Cauchy.lean +++ b/Mathlib/Topology/UniformSpace/Cauchy.lean @@ -645,6 +645,34 @@ theorem CauchySeq.totallyBounded_range {s : ℕ → α} (hs : CauchySeq s) : rcases le_total m n with hm | hm exacts [⟨m, hm, refl_mem_uniformity ha⟩, ⟨n, le_refl n, hn m hm n le_rfl⟩] +/-- Given a family of points `xs n`, a family of entourages `V n` of the diagonal and a family of +natural numbers `u n`, the intersection over `n` of the `V n`-neighborhood of `xs 1, ..., xs (u n)`. +Designed to be relatively compact when `V n` tends to the diagonal. -/ +def interUnionBalls (xs : ℕ → α) (u : ℕ → ℕ) (V : ℕ → Set (α × α)) : Set α := + ⋂ n, ⋃ m ≤ u n, UniformSpace.ball (xs m) (Prod.swap ⁻¹' V n) + +lemma totallyBounded_interUnionBalls {p : ℕ → Prop} {U : ℕ → Set (α × α)} + (H : (uniformity α).HasBasis p U) (xs : ℕ → α) (u : ℕ → ℕ) : + TotallyBounded (interUnionBalls xs u U) := by + rw [Filter.HasBasis.totallyBounded_iff H] + intro i _ + have h_subset : interUnionBalls xs u U + ⊆ ⋃ m ≤ u i, UniformSpace.ball (xs m) (Prod.swap ⁻¹' U i) := + fun x hx ↦ Set.mem_iInter.1 hx i + classical + refine ⟨Finset.image xs (Finset.range (u i + 1)), Finset.finite_toSet _, fun x hx ↦ ?_⟩ + simp only [Finset.coe_image, Finset.coe_range, mem_image, mem_Iio, iUnion_exists, biUnion_and', + iUnion_iUnion_eq_right, Nat.lt_succ_iff] + exact h_subset hx + +/-- The construction `interUnionBalls` is used to have a relatively compact set. -/ +theorem isCompact_closure_interUnionBalls {p : ℕ → Prop} {U : ℕ → Set (α × α)} + (H : (uniformity α).HasBasis p U) [CompleteSpace α] (xs : ℕ → α) (u : ℕ → ℕ) : + IsCompact (closure (interUnionBalls xs u U)) := by + rw [isCompact_iff_totallyBounded_isComplete] + refine ⟨TotallyBounded.closure ?_, isClosed_closure.isComplete⟩ + exact totallyBounded_interUnionBalls H xs u + /-! ### Sequentially complete space diff --git a/Mathlib/Topology/UniformSpace/Completion.lean b/Mathlib/Topology/UniformSpace/Completion.lean index 7e1f98e7b1de7..d7cc641b4699f 100644 --- a/Mathlib/Topology/UniformSpace/Completion.lean +++ b/Mathlib/Topology/UniformSpace/Completion.lean @@ -310,12 +310,13 @@ instance completeSpace : CompleteSpace (Completion α) := instance t0Space : T0Space (Completion α) := SeparationQuotient.instT0Space +variable {α} in /-- The map from a uniform space to its completion. -/ @[coe] def coe' : α → Completion α := SeparationQuotient.mk ∘ pureCauchy /-- Automatic coercion from `α` to its completion. Not always injective. -/ instance : Coe α (Completion α) := - ⟨coe' α⟩ + ⟨coe'⟩ -- note [use has_coe_t] protected theorem coe_eq : ((↑) : α → Completion α) = SeparationQuotient.mk ∘ pureCauchy := rfl diff --git a/Mathlib/Topology/UnitInterval.lean b/Mathlib/Topology/UnitInterval.lean index cec95ceaeefe6..5cecb700fae18 100644 --- a/Mathlib/Topology/UnitInterval.lean +++ b/Mathlib/Topology/UnitInterval.lean @@ -122,8 +122,6 @@ def symmHomeomorph : I ≃ₜ I where theorem strictAnti_symm : StrictAnti σ := fun _ _ h ↦ sub_lt_sub_left (α := ℝ) h _ -@[deprecated (since := "2024-02-27")] alias involutive_symm := symm_involutive -@[deprecated (since := "2024-02-27")] alias bijective_symm := symm_bijective @[simp] theorem symm_inj {i j : I} : σ i = σ j ↔ i = j := symm_bijective.injective.eq_iff diff --git a/Mathlib/Topology/VectorBundle/Basic.lean b/Mathlib/Topology/VectorBundle/Basic.lean index 63d31dc19d070..9ba91b204f66f 100644 --- a/Mathlib/Topology/VectorBundle/Basic.lean +++ b/Mathlib/Topology/VectorBundle/Basic.lean @@ -247,7 +247,7 @@ open Classical in def coordChangeL (e e' : Trivialization F (π F E)) [e.IsLinear R] [e'.IsLinear R] (b : B) : F ≃L[R] F := { toLinearEquiv := if hb : b ∈ e.baseSet ∩ e'.baseSet - then (e.linearEquivAt R b (hb.1 : _)).symm.trans (e'.linearEquivAt R b hb.2) + then (e.linearEquivAt R b (hb.1 :)).symm.trans (e'.linearEquivAt R b hb.2) else LinearEquiv.refl R F continuous_toFun := by by_cases hb : b ∈ e.baseSet ∩ e'.baseSet diff --git a/Mathlib/Topology/VectorBundle/Constructions.lean b/Mathlib/Topology/VectorBundle/Constructions.lean index 731ddade673df..0d6e705e2f422 100644 --- a/Mathlib/Topology/VectorBundle/Constructions.lean +++ b/Mathlib/Topology/VectorBundle/Constructions.lean @@ -145,7 +145,7 @@ theorem Trivialization.continuousLinearEquivAt_prod {e₁ : Trivialization F₁ ext v : 2 obtain ⟨v₁, v₂⟩ := v rw [(e₁.prod e₂).continuousLinearEquivAt_apply 𝕜, Trivialization.prod] - exact (congr_arg Prod.snd (prod_apply 𝕜 hx.1 hx.2 v₁ v₂) : _) + exact (congr_arg Prod.snd (prod_apply 𝕜 hx.1 hx.2 v₁ v₂) :) end diff --git a/MathlibTest/HashCommandLinter.lean b/MathlibTest/HashCommandLinter.lean index cf5304e11e734..4193342406c14 100644 --- a/MathlibTest/HashCommandLinter.lean +++ b/MathlibTest/HashCommandLinter.lean @@ -2,12 +2,6 @@ import Lean.Elab.GuardMsgs import Mathlib.Tactic.AdaptationNote import Mathlib.Tactic.Linter.HashCommandLinter - --- #adaptation_note --- The hashCommand linter started failing after https://github.com/leanprover/lean4/pull/6299 --- Disabling aync elaboration fixes it, but of course we're not going to do that globally. -set_option Elab.async false - set_option linter.hashCommand true section ignored_commands diff --git a/MathlibTest/TypeCheck.lean b/MathlibTest/TypeCheck.lean index 9d38a614053df..1345ea5d983f1 100644 --- a/MathlibTest/TypeCheck.lean +++ b/MathlibTest/TypeCheck.lean @@ -29,7 +29,7 @@ example : True := by type_check Bool.true -- Bool type_check nat_lit 1 -- Nat type_check (1 : Nat) -- Nat - type_check (True : _) -- Prop + type_check (True :) -- Prop type_check ∀ x y : Nat, x = y -- Prop type_check fun x : Nat => 2 * x + 1 -- Nat -> Nat type_check [1] diff --git a/MathlibTest/hint.lean b/MathlibTest/hint.lean index e793160ca9e7e..f2757fd22d69d 100644 --- a/MathlibTest/hint.lean +++ b/MathlibTest/hint.lean @@ -51,3 +51,26 @@ info: Try these: -/ #guard_msgs in example {P : Nat → Prop} (h : { x // P x }) : ∃ x, P x ∧ 0 ≤ x := by hint + +section multiline_hint + +local macro "this_is_a_multiline_exact" ppLine t:term : tactic => `(tactic| exact $t) + +local elab tk:"long_trivial" : tactic => do + let triv := Lean.mkIdent ``trivial + let actual ← `(tactic| this_is_a_multiline_exact $triv) + Lean.Meta.Tactic.TryThis.addSuggestion tk { suggestion := .tsyntax actual} + Lean.Elab.Tactic.evalTactic actual + +register_hint long_trivial + +/-- +info: Try these: +• this_is_a_multiline_exact + trivial +-/ +#guard_msgs in +example : True := by + hint + +end multiline_hint diff --git a/docs/100.yaml b/docs/100.yaml index 991d6a09ae840..e7f8eb83b68c0 100644 --- a/docs/100.yaml +++ b/docs/100.yaml @@ -20,9 +20,8 @@ title : Gödel’s Incompleteness Theorem author : Shogo Saito links : - results : - - First: https://github.com/FormalizedFormalLogic/Incompleteness/blob/master/Incompleteness/Arith/First.lean - - Second: https://github.com/FormalizedFormalLogic/Incompleteness/blob/master/Incompleteness/Arith/Second.lean + First incompleteness theorem: https://github.com/FormalizedFormalLogic/Incompleteness/blob/master/Incompleteness/Arith/First.lean + Second incompleteness theorem: https://github.com/FormalizedFormalLogic/Incompleteness/blob/master/Incompleteness/Arith/Second.lean website: https://formalizedformallogic.github.io/Book/ 7: title : Law of Quadratic Reciprocity diff --git a/docs/1000.yaml b/docs/1000.yaml index 50d0e3fb4af8f..b7f982e4112d6 100644 --- a/docs/1000.yaml +++ b/docs/1000.yaml @@ -39,6 +39,9 @@ Q32182: Q33481: title: Arrow's impossibility theorem + author: Benjamin Davidson, Andrew Souther + date: 2021 + url: https://github.com/asouther4/lean-social-choice/blob/master/src/arrows_theorem.lean Q98831: title: Multiplication theorem @@ -54,9 +57,13 @@ Q132427: Q132469: title: Fermat's Last Theorem + # See https://imperialcollegelondon.github.io/FLT/ Q137164: title: Besicovitch covering theorem + decl: Besicovitch.exists_disjoint_closedBall_covering_ae + author: Sébastien Gouëzel + date: 2021 Q154210: title: CPCTC @@ -72,6 +79,9 @@ Q174955: Q179208: title: Cayley's theorem + decl: Equiv.Perm.subgroupOfMulAction + author: Eric Wieser + date: 2021 Q179467: title: Fourier theorem @@ -88,12 +98,17 @@ Q180345: Q182505: title: Bayes' theorem decl: ProbabilityTheory.cond_eq_inv_mul_cond_mul + author: Rishikesh Vaishnav + date: 2022 Q184410: title: Four color theorem Q188295: title: Fermat's little theorem + decl: ZMod.pow_card_sub_one_eq_one + author: Aaron Anderson + date: 2020 Q188745: title: Classification of Platonic solids @@ -104,6 +119,7 @@ Q189136: # XXX: other branch found exists_ratio_deriv_eq_ratio_slope, # is about Cauchy's vs Lagrange's Mean Value Theorem, TODO decide! author: Yury G. Kudryashov + date: 2019 Q190026: title: Ford's theorem @@ -119,16 +135,25 @@ Q190556: title: De Moivre's theorem decl: Complex.cos_add_sin_mul_I_pow author: Abhimanyu Pallavi Sudhir + date: 2019 Q191693: title: Lebesgue's decomposition theorem + decl: MeasureTheory.Measure.haveLebesgueDecomposition_of_sigmaFinite + author: Kexing Ying + date: 2021 Q192760: title: Fundamental theorem of algebra + decl: Complex.isAlgClosed + author: Chris Hughes + date: 2019 Q193286: title: Rolle's theorem decl: exists_deriv_eq_zero + author: Yury G. Kudryashov + date: 2019 Q193878: title: Chinese remainder theorem @@ -159,6 +184,7 @@ Q203565: title: Solutions of a general cubic equation decl: Theorems100.cubic_eq_zero_iff author: Jeoff Lee + date: 2022 Q204884: title: Löb's theorem @@ -529,6 +555,9 @@ Q608294: Q609612: title: Knaster–Tarski theorem + decl: fixedPoints.completeLattice + author: Kenny Lau + date: 2018 Q612021: title: Gibbard–Satterthwaite theorem @@ -563,6 +592,7 @@ Q642620: title: Krein–Milman theorem decl: closure_convexHull_extremePoints author: Yaël Dillies + date: 2022 Q643513: title: Riesz–Fischer theorem @@ -655,15 +685,22 @@ Q716171: title: Erdős–Ginzburg–Ziv theorem decl: ZMod.erdos_ginzburg_ziv author: Yaël Dillies + date: 2023 Q718875: title: Erdős–Ko–Rado theorem + decl: Finset.erdos_ko_rado + author: Bhavik Mehta, Yaël Dillies + date: 2020 Q719966: title: Toda's theorem Q720469: title: Chevalley–Warning theorem + decl: char_dvd_card_solutions_of_sum_lt + author: Johan Commelin + date: 2019 Q721695: title: Szemerédi–Trotter theorem @@ -702,6 +739,9 @@ Q744440: Q748233: title: Sylvester–Gallai theorem + author: Bhavik Mehta + url: https://github.com/YaelDillies/LeanCamCombi/blob/6a6a670f324b2af82ae17f21f9d51ac0bc859f6f/LeanCamCombi/SylvesterChvatal.lean#L610 + date: 2022 Q751120: title: Thue–Siegel–Roth theorem @@ -709,6 +749,8 @@ Q751120: Q752375: title: Extreme value theorem decl: IsCompact.exists_isMinOn + author: Sébastien Gouëzel + date: 2018 Q755986: title: Atiyah–Bott fixed-point theorem @@ -720,6 +762,7 @@ Q756946: title: Lagrange's four-square theorem decl: Nat.sum_four_squares author: Chris Hughes + date: 2019 Q764287: title: Van der Waerden's theorem @@ -777,6 +820,9 @@ Q830513: Q834025: title: Cauchy integral theorem + decl: Complex.circleIntegral_div_sub_of_differentiable_on_off_countable + author: Yury Kudryashov + date: 2021 Q834211: title: Wallace–Bolyai–Gerwien theorem @@ -793,6 +839,9 @@ Q841893: Q842953: title: Lebesgue's density theorem + decl: Besicovitch.ae_tendsto_measure_inter_div + author: Oliver Nash + date: 2022 Q844612: title: Brun's theorem @@ -967,8 +1016,10 @@ Q944297: # Banach open mapping theorem/Banach-Schauder theorem in functional analysis is in mathlib (ContinuousLinearMap.isOpenMap) Q948664: - title: Kneser's theorem - # TODO: is this in mathlib already? + title: Kneser's addition theorem + author: Mantas Bakšys, Yaël Dillies + url: https://github.com/YaelDillies/LeanCamCombi/blob/master/LeanCamCombi/Kneser/Kneser.lean + date: 2022 Q951327: title: Pasch's theorem @@ -1166,6 +1217,7 @@ Q1097021: title: Minkowski's theorem decl: MeasureTheory.exists_ne_zero_mem_lattice_of_measure_mul_two_pow_lt_measure author: Alex J. Best and Yaël Dillies + date: 2021 Q1103054: title: Lions–Lax–Milgram theorem @@ -1364,6 +1416,9 @@ Q1242398: Q1243340: title: Birkhoff–Von Neumann theorem + decl: doublyStochastic_eq_convexHull_permMatrix + author: Bhavik Mehta + date: 2024 Q1249069: title: Richardson's theorem @@ -1471,6 +1526,8 @@ Q1425529: Q1426292: title: Banach–Steinhaus theorem decl: banach_steinhaus + author: Jireh Loreaux + date: 2021 Q1434158: title: Fluctuation dissipation theorem @@ -1551,6 +1608,8 @@ Q1568612: Q1568811: title: Hahn decomposition theorem decl: MeasureTheory.SignedMeasure.exists_isCompl_positive_negative + author: Kexing Ying + date: 2021 Q1572474: title: Lindemann–Weierstrass theorem @@ -1582,6 +1641,9 @@ Q1632301: Q1632433: title: Helly's theorem + author: Vasily Nesterov + decl: Convex.helly_theorem + date: 2023 Q1637085: title: Poynting's theorem @@ -1597,6 +1659,9 @@ Q1683356: Q1687147: title: Sprague–Grundy theorem + decl: SetTheory.PGame.equiv_nim_grundyValue + author: Fox Thomson + date: 2020 Q1694565: title: Vinogradov's theorem @@ -1733,6 +1798,9 @@ Q2022775: Q2027347: title: Optional stopping theorem + decl: MeasureTheory.submartingale_iff_expected_stoppedValue_mono + author: Kexing Ying, Rémy Degenne + date: 2022 Q2028341: title: Ado's theorem @@ -1825,10 +1893,15 @@ Q2226774: Q2226786: title: Sperner's theorem + decl: IsAntichain.sperner + author: Bhavik Mehta, Alena Gusakov, Yaël Dillies + date: 2022 Q2226800: title: Schur–Zassenhaus theorem decl: Subgroup.exists_left_complement'_of_coprime + author: Thomas Browning + date: 2021 Q2226803: title: Tennenbaum's theorem @@ -1837,7 +1910,10 @@ Q2226807: title: Vantieghems theorem Q2226855: - title: Sarkovskii's theorem + title: Sharkovskii's theorem + # TODO: Find the code + author: Bhavik Mehta + date: 2022 Q2226868: title: Geometric mean theorem @@ -1864,6 +1940,9 @@ Q2270905: Q2275191: title: Lebesgue differentiation theorem + decl: VitaliFamily.ae_tendsto_lintegral_div + author: Sébastien Gouëzel + date: 2021 Q2277040: title: Shannon's expansion theorem @@ -1882,6 +1961,7 @@ Q2310718: Q2338929: title: Exchange theorem + # aka Steinitz exchange lemma. This must definitely be in mathlib somewhere Q2345282: title: Shannon's theorem @@ -1933,6 +2013,7 @@ Q2478371: Q2482753: title: Mann's theorem + # Done in LeanCamCombi Q2495664: title: Universal coefficient theorem @@ -1951,6 +2032,7 @@ Q2524990: Q2525646: title: Jordan–Hölder theorem + decl: CompositionSeries.jordan_holder Q2533936: title: Descartes's theorem on total angular defect @@ -2011,6 +2093,9 @@ Q2916568: Q2919401: title: Ostrowski's theorem + decl: Rat.AbsoluteValue.equiv_real_or_padic + author: David Kurniadi Angdinata, Fabrizio Barroero, Laura Capuano, Nirvana Coppola, María Inés de Frutos-Fernández, Sam van Gool, Silvain Rideau-Kikuchi, Amos Turchet, Francesco Veneziano + date: 2024 Q2981012: title: Monge's theorem @@ -2056,6 +2141,9 @@ Q3229335: Q3229352: title: Vitali covering theorem + decl: Vitali.exists_disjoint_covering_ae + author: Sébastien Gouëzel + date: 2021 Q3345678: title: Moore–Aronszajn theorem @@ -2092,6 +2180,9 @@ Q3526993: Q3526996: title: Kolmogorov extension theorem + author: Rémy Degenne, Peter Pfaffelhuber + date: 2023 + url: https://github.com/RemyDegenne/kolmogorov_extension4 Q3526998: title: Excision theorem @@ -2161,6 +2252,9 @@ Q3527100: Q3527102: title: Kruskal–Katona theorem + decl: Finset.kruskal_katona + author: Bhavik Mehta, Yaël Dillies + date: 2020 Q3527110: title: Lax–Wendroff theorem @@ -2212,6 +2306,9 @@ Q3527214: Q3527215: title: Hilbert projection theorem + decl: exists_norm_eq_iInf_of_complete_convex + author: Zhouhang Zhou + date: 2019 Q3527217: title: M. Riesz extension theorem @@ -2242,9 +2339,15 @@ Q3527247: Q3527250: title: Hadamard three-lines theorem + decl: Complex.HadamardThreeLines.norm_le_interpStrip_of_mem_verticalClosedStrip + author: Xavier Généreux + date: 2023 Q3527263: title: Kleene fixed-point theorem + decl: fixedPoints.lfp_eq_sSup_iterate + author: Ira Fesefeldt + date: 2024 Q3527279: title: Wiener's tauberian theorem @@ -2299,6 +2402,9 @@ Q3984052: Q3984053: title: Fourier inversion theorem + decl: Continuous.fourier_inversion + author: Sébastien Gouëzel + date: 2024 Q3984056: title: No-communication theorem @@ -2329,6 +2435,9 @@ Q4272645: Q4378868: title: Phragmén–Lindelöf theorem + decl: PhragmenLindelof.horizontal_strip + author: Yury G. Kudryashov + date: 2022 Q4378889: title: Carathéodory's theorem @@ -2416,6 +2525,9 @@ Q4677985: Q4695203: title: Four functions theorem + decl: four_functions_theorem + author: Yaël Dillies + date: 2023 Q4700718: title: Akhiezer's theorem @@ -2485,6 +2597,9 @@ Q4827308: Q4830725: title: Ax–Grothendieck theorem + decl: ax_grothendieck_zeroLocus + author: Chris Hughes + date: 2023 Q4832965: title: Aztec diamond theorem @@ -2698,6 +2813,9 @@ Q5166389: Q5171652: title: Corners theorem + decl: corners_theorem_nat + author: Yaël Dillies, Bhavik Mehta + date: 2022 Q5172143: title: Corona theorem @@ -3366,6 +3484,9 @@ Q7433034: Q7433182: title: Schwartz–Zippel theorem + decl: MvPolynomial.schwartz_zippel_sup_sum + author: Bolton Bailey, Yaël Dillies + date: 2023 Q7433295: title: Osterwalder–Schrader theorem @@ -3524,6 +3645,9 @@ Q7857350: Q7888360: title: Structure theorem for finitely generated modules over a principal ideal domain + decl: Module.equiv_free_prod_directSum + author: Pierre-Alexandre Bazin + date: 2022 Q7894110: title: Universal approximation theorem @@ -3587,6 +3711,9 @@ Q8066795: Q8074796: title: Zsigmondy's theorem + # TODO: Find the code + author: Mantas Bakšys + date: 2022 Q8081891: title: Śleszyński–Pringsheim theorem @@ -3605,6 +3732,9 @@ Q10942247: Q11352023: title: Vitali convergence theorem + decl: MeasureTheory.tendstoInMeasure_iff_tendsto_Lp + author: Igor Khavkine + date: 2024 Q11573495: title: Plancherel theorem for spherical functions @@ -3641,6 +3771,9 @@ Q15895894: Q16251580: title: Matiyasevich's theorem + decl: Pell.matiyasevic + author: Mario Carneiro + date: 2017 Q16680059: title: Strong perfect graph theorem @@ -3659,6 +3792,9 @@ Q17003552: Q17005116: title: Birkhoff's representation theorem + decl: LatticeHom.birkhoffSet + author: Yaël Dillies + date: 2022 Q17008559: title: Davenport–Schmidt theorem diff --git a/docs/references.bib b/docs/references.bib index 20343a1c0a28c..726943e64c4fe 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -2521,6 +2521,18 @@ @Book{ lam_1999 year = {1999} } +@Article{ Lambek_1964, + title = {A Module is Flat if and Only if its Character Module is + Injective}, + volume = {7}, + doi = {10.4153/CMB-1964-021-9}, + number = {2}, + journal = {Canadian Mathematical Bulletin}, + author = {Lambek, J.}, + year = {1964}, + pages = {237–243} +} + @Article{ lazarus1973, author = {Michel Lazarus}, title = {Les familles libres maximales d'un module ont-elles le diff --git a/lakefile.lean b/lakefile.lean index a9a535eb971a2..8547cccd1b924 100644 --- a/lakefile.lean +++ b/lakefile.lean @@ -24,7 +24,8 @@ require "leanprover-community" / "plausible" @ git "v4.16.0-rc1" * as `moreServerArgs`, to set their default value in mathlib (as well as `Archive`, `Counterexamples` and `test`). -/ abbrev mathlibOnlyLinters : Array LeanOption := #[ - ⟨`linter.docPrime, true⟩, + -- The `docPrime` linter is disabled: https://github.com/leanprover-community/mathlib4/issues/20560 + ⟨`linter.docPrime, false⟩, ⟨`linter.hashCommand, true⟩, ⟨`linter.oldObtain, true,⟩, ⟨`linter.refine, true⟩, diff --git a/scripts/bench_summary.lean b/scripts/bench_summary.lean index 9491ddef2ee99..12d9ec7fe769b 100644 --- a/scripts/bench_summary.lean +++ b/scripts/bench_summary.lean @@ -147,7 +147,9 @@ open Lean Elab Command in as a comment to a pull request. It takes as input * the number `PR` and the name `repo` as a `String` containing the relevant pull-request (it reads and posts comments there) -* the optional `jobID` string for reporting the action that produced the output +* the optional `jobID` numeral for reporting the action that produced the output + (`jobID` is a natural number, even though it gets converted to a `String` -- this is mostly + due to the fact that it is easier to get CI to pass a number, than a string with quotations) * the `String` `tempFile` of a temporary file where the command stores transient information. The code itself interfaces with the shell to retrieve and process json data and eventually @@ -163,7 +165,7 @@ Here is a summary of the steps: * process the final string to produce a summary (using `benchOutput`), * finally post the resulting output to the PR (using `gh pr comment ...`). -/ -def addBenchSummaryComment (PR : Nat) (repo : String) (jobID : String := "") +def addBenchSummaryComment (PR : Nat) (repo : String) (jobID : Nat := 0) (author : String := "leanprover-bot") (tempFile : String := "benchOutput.json") : CommandElabM Unit := do let job_msg := s!"\n[CI run](https://github.com/{repo}/actions/runs/{jobID})" diff --git a/scripts/technical-debt-metrics.sh b/scripts/technical-debt-metrics.sh index 7c982cadbcb20..efe28e4c7e85d 100755 --- a/scripts/technical-debt-metrics.sh +++ b/scripts/technical-debt-metrics.sh @@ -80,7 +80,6 @@ titlesPathsAndRegexes=( "adaptation notes" "*" "adaptation_note" "disabled simpNF lints" "*" "nolint simpNF" "erw" "*" "erw \[" - "ofNat" "*" "See note \[no_index around OfNat.ofNat\]" "maxHeartBeats modifications" ":^MathlibTest" "^ *set_option .*maxHeartbeats" ) @@ -160,12 +159,14 @@ then then printf '
No changes to technical debt.\n' else - printf '%s\n' "${rep}" | + printf '%s\n' "${rep}" | # outputs lines containing `|Current number|Change|Type|`, so + # `$2` refers to `Current number` and `$3` to `Change`. awk -F '|' -v rep="${rep}" ' - BEGIN{total=0; weight=0} + BEGIN{total=0; weight=0; absWeight=0} + {absWeight+=$3+0} (($3+0 == $3) && (!($2+0 == 0))) {total+=1 / $2; weight+=$3 / $2} END{ - average=weight/total + if (total == 0) {average=absWeight} else {average=weight/total} if(average < 0) {change= "Decrease"; average=-average; weight=-weight} else {change= "Increase"} printf("
%s in tech debt: (relative, absolute) = (%4.2f, %4.2f)\n\n%s\n", change, average, weight, rep) }' fi