From 54d30b77df91e4f105dcccef562c059aca732b85 Mon Sep 17 00:00:00 2001 From: qawbecrdtey Date: Sun, 6 Oct 2024 19:08:01 +0000 Subject: [PATCH 01/43] feat: add `Finset.cons_swap` (#17413) ``` (s.cons a _).cons b _ = (s.cons b _).cons a _ ``` Co-authored-by: qawbecrdtey Co-authored-by: qawbecrdtey <40463813+qawbecrdtey@users.noreply.github.com> --- Mathlib/Data/Finset/Basic.lean | 7 ++++++- 1 file changed, 6 insertions(+), 1 deletion(-) diff --git a/Mathlib/Data/Finset/Basic.lean b/Mathlib/Data/Finset/Basic.lean index 3e1b0ec9fd4a7..e86ab424aa5fb 100644 --- a/Mathlib/Data/Finset/Basic.lean +++ b/Mathlib/Data/Finset/Basic.lean @@ -807,6 +807,11 @@ theorem ssubset_iff_exists_cons_subset : s ⊂ t ↔ ∃ (a : _) (h : a ∉ s), obtain ⟨a, hs, ht⟩ := not_subset.1 h.2 exact ⟨a, ht, cons_subset.2 ⟨hs, h.subset⟩⟩ +theorem cons_swap (hb : b ∉ s) (ha : a ∉ s.cons b hb) : + (s.cons b hb).cons a ha = (s.cons a fun h ↦ ha (mem_cons.mpr (.inr h))).cons b fun h ↦ + ha (mem_cons.mpr (.inl ((mem_cons.mp h).elim symm (fun h ↦ False.elim (hb h))))) := + eq_of_veq <| Multiset.cons_swap a b s.val + end Cons /-! ### disjoint -/ @@ -3093,4 +3098,4 @@ def proveFinsetNonempty {u : Level} {α : Q(Type u)} (s : Q(Finset $α)) : end Mathlib.Meta -set_option linter.style.longFile 3100 +set_option linter.style.longFile 3200 From 3d637cea4a4e45f51bc88a2f09ef38fd72618fb7 Mon Sep 17 00:00:00 2001 From: Moritz Firsching Date: Sun, 6 Oct 2024 20:06:17 +0000 Subject: [PATCH 02/43] chore(Order/Heyting/Hom): remove left-over TODO (#17475) This is already present. Co-authored-by: Moritz Firsching --- Mathlib/Order/Heyting/Hom.lean | 1 - 1 file changed, 1 deletion(-) diff --git a/Mathlib/Order/Heyting/Hom.lean b/Mathlib/Order/Heyting/Hom.lean index 358a7b6215806..6b0940817f6d3 100644 --- a/Mathlib/Order/Heyting/Hom.lean +++ b/Mathlib/Order/Heyting/Hom.lean @@ -190,7 +190,6 @@ theorem map_compl (a : α) : f aᶜ = (f a)ᶜ := by rw [← himp_bot, ← himp_ @[simp] theorem map_bihimp (a b : α) : f (a ⇔ b) = f a ⇔ f b := by simp_rw [bihimp, map_inf, map_himp] --- TODO: `map_bihimp` end HeytingAlgebra section CoheytingAlgebra From e7484e60aeacc624a2ae169d88d24c5ca58a5567 Mon Sep 17 00:00:00 2001 From: Markus Himmel Date: Mon, 7 Oct 2024 04:38:10 +0000 Subject: [PATCH 03/43] feat(CategoryTheory): `map_yonedaEquiv` (#17438) --- Mathlib/CategoryTheory/Yoneda.lean | 20 ++++++++++++++++++++ 1 file changed, 20 insertions(+) diff --git a/Mathlib/CategoryTheory/Yoneda.lean b/Mathlib/CategoryTheory/Yoneda.lean index 5fbeccb9b2ec8..f06afc7170c1c 100644 --- a/Mathlib/CategoryTheory/Yoneda.lean +++ b/Mathlib/CategoryTheory/Yoneda.lean @@ -417,6 +417,7 @@ theorem yonedaEquiv_symm_app_apply {X : C} {F : Cᵒᵖ ⥤ Type v₁} (x : F.ob (f : Y.unop ⟶ X) : (yonedaEquiv.symm x).app Y f = F.map f.op x := rfl +/-- See also `yonedaEquiv_naturality'` for a more general version. -/ lemma yonedaEquiv_naturality {X Y : C} {F : Cᵒᵖ ⥤ Type v₁} (f : yoneda.obj X ⟶ F) (g : Y ⟶ X) : F.map g.op (yonedaEquiv f) = yonedaEquiv (yoneda.map g ≫ f) := by change (f.app (op X) ≫ F.map g.op) (𝟙 X) = f.app (op Y) (𝟙 Y ≫ g) @@ -424,6 +425,9 @@ lemma yonedaEquiv_naturality {X Y : C} {F : Cᵒᵖ ⥤ Type v₁} (f : yoneda.o dsimp simp +/-- Variant of `yonedaEquiv_naturality` with general `g`. This is technically strictly more general + than `yonedaEquiv_naturality`, but `yonedaEquiv_naturality` is sometimes preferable because it + can avoid the "motive is not type correct" error. -/ lemma yonedaEquiv_naturality' {X Y : Cᵒᵖ} {F : Cᵒᵖ ⥤ Type v₁} (f : yoneda.obj (unop X) ⟶ F) (g : X ⟶ Y) : F.map g (yonedaEquiv f) = yonedaEquiv (yoneda.map g.unop ≫ f) := yonedaEquiv_naturality _ _ @@ -436,6 +440,18 @@ lemma yonedaEquiv_yoneda_map {X Y : C} (f : X ⟶ Y) : yonedaEquiv (yoneda.map f rw [yonedaEquiv_apply] simp +/-- See also `map_yonedaEquiv'` for a more general version. -/ +lemma map_yonedaEquiv {X Y : C} {F : Cᵒᵖ ⥤ Type v₁} (f : yoneda.obj X ⟶ F) + (g : Y ⟶ X) : F.map g.op (yonedaEquiv f) = f.app (op Y) g := by + rw [yonedaEquiv_naturality, yonedaEquiv_comp, yonedaEquiv_yoneda_map] + +/-- Variant of `map_yonedaEquiv` with general `g`. This is technically strictly more general + than `map_yonedaEquiv`, but `map_yonedaEquiv` is sometimes preferable because it + can avoid the "motive is not type correct" error. -/ +lemma map_yonedaEquiv' {X Y : Cᵒᵖ} {F : Cᵒᵖ ⥤ Type v₁} (f : yoneda.obj (unop X) ⟶ F) + (g : X ⟶ Y) : F.map g (yonedaEquiv f) = f.app Y g.unop := by + rw [yonedaEquiv_naturality', yonedaEquiv_comp, yonedaEquiv_yoneda_map] + lemma yonedaEquiv_symm_map {X Y : Cᵒᵖ} (f : X ⟶ Y) {F : Cᵒᵖ ⥤ Type v₁} (t : F.obj X) : yonedaEquiv.symm (F.map f t) = yoneda.map f.unop ≫ yonedaEquiv.symm t := by obtain ⟨u, rfl⟩ := yonedaEquiv.surjective t @@ -612,6 +628,10 @@ lemma coyonedaEquiv_coyoneda_map {X Y : C} (f : X ⟶ Y) : rw [coyonedaEquiv_apply] simp +lemma map_coyonedaEquiv {X Y : C} {F : C ⥤ Type v₁} (f : coyoneda.obj (op X) ⟶ F) + (g : X ⟶ Y) : F.map g (coyonedaEquiv f) = f.app Y g := by + rw [coyonedaEquiv_naturality, coyonedaEquiv_comp, coyonedaEquiv_coyoneda_map] + lemma coyonedaEquiv_symm_map {X Y : C} (f : X ⟶ Y) {F : C ⥤ Type v₁} (t : F.obj X) : coyonedaEquiv.symm (F.map f t) = coyoneda.map f.op ≫ coyonedaEquiv.symm t := by obtain ⟨u, rfl⟩ := coyonedaEquiv.surjective t From a014dc63b285fdb74a6c27f789f5de6441c03296 Mon Sep 17 00:00:00 2001 From: Markus Himmel Date: Mon, 7 Oct 2024 05:52:39 +0000 Subject: [PATCH 04/43] chore(CategoryTheory): generalize instances for limits of bifunctors (#17439) Also add a lemma about evaluating `colimit.map` into an element included into a colimit of a `Type`-valued bifunctor. --- .../Limits/FunctorCategory/Basic.lean | 57 ++++++++++++------- .../CategoryTheory/Limits/FunctorToTypes.lean | 13 +++-- 2 files changed, 46 insertions(+), 24 deletions(-) diff --git a/Mathlib/CategoryTheory/Limits/FunctorCategory/Basic.lean b/Mathlib/CategoryTheory/Limits/FunctorCategory/Basic.lean index 305a3bbd5a648..0b2e0d72e2f6e 100644 --- a/Mathlib/CategoryTheory/Limits/FunctorCategory/Basic.lean +++ b/Mathlib/CategoryTheory/Limits/FunctorCategory/Basic.lean @@ -146,18 +146,23 @@ def combinedIsColimit (F : J ⥤ K ⥤ C) (c : ∀ k : K, ColimitCocone (F.flip. noncomputable section +instance functorCategoryHasLimit (F : J ⥤ K ⥤ C) [∀ k, HasLimit (F.flip.obj k)] : HasLimit F := + HasLimit.mk + { cone := combineCones F fun _ => getLimitCone _ + isLimit := combinedIsLimit _ _ } + instance functorCategoryHasLimitsOfShape [HasLimitsOfShape J C] : HasLimitsOfShape J (K ⥤ C) where - has_limit F := - HasLimit.mk - { cone := combineCones F fun _ => getLimitCone _ - isLimit := combinedIsLimit _ _ } + has_limit _ := inferInstance + +instance functorCategoryHasColimit (F : J ⥤ K ⥤ C) [∀ k, HasColimit (F.flip.obj k)] : + HasColimit F := + HasColimit.mk + { cocone := combineCocones F fun _ => getColimitCocone _ + isColimit := combinedIsColimit _ _ } instance functorCategoryHasColimitsOfShape [HasColimitsOfShape J C] : HasColimitsOfShape J (K ⥤ C) where - has_colimit _ := - HasColimit.mk - { cocone := combineCocones _ fun _ => getColimitCocone _ - isColimit := combinedIsColimit _ _ } + has_colimit _ := inferInstance -- Porting note: previously Lean could see through the binders and infer_instance sufficed instance functorCategoryHasLimitsOfSize [HasLimitsOfSize.{v₁, u₁} C] : @@ -169,14 +174,20 @@ instance functorCategoryHasColimitsOfSize [HasColimitsOfSize.{v₁, u₁} C] : HasColimitsOfSize.{v₁, u₁} (K ⥤ C) where has_colimits_of_shape := fun _ _ => inferInstance +instance hasLimitCompEvalution (F : J ⥤ K ⥤ C) (k : K) [HasLimit (F.flip.obj k)] : + HasLimit (F ⋙ (evaluation _ _).obj k) := + hasLimitOfIso (F := F.flip.obj k) (Iso.refl _) + +instance evaluationPreservesLimit (F : J ⥤ K ⥤ C) [∀ k, HasLimit (F.flip.obj k)] (k : K) : + PreservesLimit F ((evaluation K C).obj k) := + -- Porting note: added a let because X was not inferred + let X : (k : K) → LimitCone (F.flip.obj k) := fun k => getLimitCone (F.flip.obj k) + preservesLimitOfPreservesLimitCone (combinedIsLimit _ X) <| + IsLimit.ofIsoLimit (limit.isLimit _) (evaluateCombinedCones F X k).symm + instance evaluationPreservesLimitsOfShape [HasLimitsOfShape J C] (k : K) : PreservesLimitsOfShape J ((evaluation K C).obj k) where - preservesLimit {F} := by - -- Porting note: added a let because X was not inferred - let X : (k : K) → LimitCone (Prefunctor.obj (Functor.flip F).toPrefunctor k) := - fun k => getLimitCone (Prefunctor.obj (Functor.flip F).toPrefunctor k) - exact preservesLimitOfPreservesLimitCone (combinedIsLimit _ _) <| - IsLimit.ofIsoLimit (limit.isLimit _) (evaluateCombinedCones F X k).symm + preservesLimit := inferInstance /-- If `F : J ⥤ K ⥤ C` is a functor into a functor category which has a limit, then the evaluation of that limit at `k` is the limit of the evaluations of `F.obj j` at `k`. @@ -225,14 +236,20 @@ theorem limit_obj_ext {H : J ⥤ K ⥤ C} [HasLimitsOfShape J C] {k : K} {W : C} ext j simpa using w j +instance hasColimitCompEvaluation (F : J ⥤ K ⥤ C) (k : K) [HasColimit (F.flip.obj k)] : + HasColimit (F ⋙ (evaluation _ _).obj k) := + hasColimitOfIso (F := F.flip.obj k) (Iso.refl _) + +instance evaluationPreservesColimit (F : J ⥤ K ⥤ C) [∀ k, HasColimit (F.flip.obj k)] (k : K) : + PreservesColimit F ((evaluation K C).obj k) := + -- Porting note: added a let because X was not inferred + let X : (k : K) → ColimitCocone (F.flip.obj k) := fun k => getColimitCocone (F.flip.obj k) + preservesColimitOfPreservesColimitCocone (combinedIsColimit _ X) <| + IsColimit.ofIsoColimit (colimit.isColimit _) (evaluateCombinedCocones F X k).symm + instance evaluationPreservesColimitsOfShape [HasColimitsOfShape J C] (k : K) : PreservesColimitsOfShape J ((evaluation K C).obj k) where - preservesColimit {F} := by - -- Porting note: added a let because X was not inferred - let X : (k : K) → ColimitCocone (Prefunctor.obj (Functor.flip F).toPrefunctor k) := - fun k => getColimitCocone (Prefunctor.obj (Functor.flip F).toPrefunctor k) - refine preservesColimitOfPreservesColimitCocone (combinedIsColimit _ _) <| - IsColimit.ofIsoColimit (colimit.isColimit _) (evaluateCombinedCocones F X k).symm + preservesColimit := inferInstance /-- If `F : J ⥤ K ⥤ C` is a functor into a functor category which has a colimit, then the evaluation of that colimit at `k` is the colimit of the evaluations of `F.obj j` at `k`. diff --git a/Mathlib/CategoryTheory/Limits/FunctorToTypes.lean b/Mathlib/CategoryTheory/Limits/FunctorToTypes.lean index 62eb96c6be392..ba9abc78c25d6 100644 --- a/Mathlib/CategoryTheory/Limits/FunctorToTypes.lean +++ b/Mathlib/CategoryTheory/Limits/FunctorToTypes.lean @@ -20,15 +20,20 @@ open CategoryTheory.Limits universe w v₁ v₂ u₁ u₂ variable {J : Type u₁} [Category.{v₁} J] {K : Type u₂} [Category.{v₂} K] -variable (F : J ⥤ K ⥤ TypeMax.{u₁, w}) +variable (F : J ⥤ K ⥤ Type w) -theorem jointly_surjective (k : K) {t : Cocone F} (h : IsColimit t) (x : t.pt.obj k) : - ∃ j y, x = (t.ι.app j).app k y := by +theorem jointly_surjective (k : K) {t : Cocone F} (h : IsColimit t) (x : t.pt.obj k) + [∀ k, HasColimit (F.flip.obj k)] : ∃ j y, x = (t.ι.app j).app k y := by let hev := isColimitOfPreserves ((evaluation _ _).obj k) h obtain ⟨j, y, rfl⟩ := Types.jointly_surjective _ hev x exact ⟨j, y, by simp⟩ -theorem jointly_surjective' (k : K) (x : (colimit F).obj k) : ∃ j y, x = (colimit.ι F j).app k y := +theorem jointly_surjective' [∀ k, HasColimit (F.flip.obj k)] (k : K) (x : (colimit F).obj k) : + ∃ j y, x = (colimit.ι F j).app k y := jointly_surjective _ _ (colimit.isColimit _) x +theorem colimit.map_ι_apply [HasColimit F] (j : J) {k k' : K} {f : k ⟶ k'} {x} : + (colimit F).map f ((colimit.ι F j).app _ x) = (colimit.ι F j).app _ ((F.obj j).map f x) := + congrFun ((colimit.ι F j).naturality _).symm _ + end CategoryTheory.FunctorToTypes From 0ae05318587d24736e49307e9e1230be592799f7 Mon Sep 17 00:00:00 2001 From: Oliver Nash Date: Mon, 7 Oct 2024 05:52:40 +0000 Subject: [PATCH 05/43] feat: a reflexive vector space (in the purely algebraic sense) is finite dimensional (#17456) --- Mathlib/LinearAlgebra/Dual.lean | 17 +++++++++++++++++ 1 file changed, 17 insertions(+) diff --git a/Mathlib/LinearAlgebra/Dual.lean b/Mathlib/LinearAlgebra/Dual.lean index 889c5e190d266..734efc1dd1911 100644 --- a/Mathlib/LinearAlgebra/Dual.lean +++ b/Mathlib/LinearAlgebra/Dual.lean @@ -580,6 +580,7 @@ class IsReflexive : Prop where lemma bijective_dual_eval [IsReflexive R M] : Bijective (Dual.eval R M) := IsReflexive.bijective_dual_eval' +/-- See also `Module.instFiniteDimensionalOfIsReflexive` for the converse over a field. -/ instance IsReflexive.of_finite_of_free [Module.Finite R M] [Free R M] : IsReflexive R M where bijective_dual_eval' := ⟨LinearMap.ker_eq_bot.mp (Free.chooseBasis R M).eval_ker, LinearMap.range_eq_top.mp (Free.chooseBasis R M).eval_range⟩ @@ -650,6 +651,22 @@ instance _root_.MulOpposite.instModuleIsReflexive : IsReflexive R (MulOpposite M instance _root_.ULift.instModuleIsReflexive.{w} : IsReflexive R (ULift.{w} M) := equiv ULift.moduleEquiv.symm +instance instFiniteDimensionalOfIsReflexive (K V : Type*) + [Field K] [AddCommGroup V] [Module K V] [IsReflexive K V] : + FiniteDimensional K V := by + rw [FiniteDimensional, ← rank_lt_aleph0_iff] + by_contra! contra + suffices lift (Module.rank K V) < Module.rank K (Dual K (Dual K V)) by + have heq := lift_rank_eq_of_equiv_equiv (R := K) (R' := K) (M := V) (M' := Dual K (Dual K V)) + (ZeroHom.id K) (evalEquiv K V) bijective_id (fun r v ↦ (evalEquiv K V).map_smul _ _) + rw [← lift_umax, heq, lift_id'] at this + exact lt_irrefl _ this + have h₁ : lift (Module.rank K V) < Module.rank K (Dual K V) := lift_rank_lt_rank_dual contra + have h₂ : Module.rank K (Dual K V) < Module.rank K (Dual K (Dual K V)) := by + convert lift_rank_lt_rank_dual <| le_trans (by simpa) h₁.le + rw [lift_id'] + exact lt_trans h₁ h₂ + end IsReflexive end Module From 3ee779f092be35426a5d7f7ed8e3489014e2ed4d Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Mon, 7 Oct 2024 06:54:57 +0000 Subject: [PATCH 06/43] chore(SetTheory/Ordinal/Basic): remove simps from `enumIsoToType` (#16907) Since lemmas like [`Ordinal.enum_lt_enum`](https://leanprover-community.github.io/mathlib4_docs/Mathlib/SetTheory/Ordinal/Basic.html#Ordinal.enum_lt_enum) can't be made into `simp` lemmas, we actually *lose* `simp` capability by having the `simps` attribute here. This was breaking some proofs in my nimber project. --- Mathlib/SetTheory/Ordinal/Basic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/SetTheory/Ordinal/Basic.lean b/Mathlib/SetTheory/Ordinal/Basic.lean index 895d453ab84d9..ffe4d081350be 100644 --- a/Mathlib/SetTheory/Ordinal/Basic.lean +++ b/Mathlib/SetTheory/Ordinal/Basic.lean @@ -1002,7 +1002,7 @@ theorem enum_inj {r : α → α → Prop} [IsWellOrder α r] {o₁ o₂ : {o // rw [EmbeddingLike.apply_eq_iff_eq, Subtype.mk.injEq] /-- The order isomorphism between ordinals less than `o` and `o.toType`. -/ -@[simps!] +@[simps! (config := .lemmasOnly)] noncomputable def enumIsoToType (o : Ordinal) : Set.Iio o ≃o o.toType where toFun x := enum (α := o.toType) (· < ·) ⟨x.1, by From e260f259d7cbe54699643b264d6dab483b809662 Mon Sep 17 00:00:00 2001 From: Markus Himmel Date: Mon, 7 Oct 2024 08:00:09 +0000 Subject: [PATCH 07/43] feat(CategoryTheory): variants of `Functor.flipCompEvaluation` (#17407) Co-authored-by: Markus Himmel --- Mathlib/CategoryTheory/Functor/Currying.lean | 1 + Mathlib/CategoryTheory/Monoidal/Category.lean | 1 + Mathlib/CategoryTheory/Products/Basic.lean | 14 ++++++++++++-- 3 files changed, 14 insertions(+), 2 deletions(-) diff --git a/Mathlib/CategoryTheory/Functor/Currying.lean b/Mathlib/CategoryTheory/Functor/Currying.lean index edaba6e369347..e130aedb422c8 100644 --- a/Mathlib/CategoryTheory/Functor/Currying.lean +++ b/Mathlib/CategoryTheory/Functor/Currying.lean @@ -3,6 +3,7 @@ Copyright (c) 2017 Kim Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison -/ +import Mathlib.CategoryTheory.EqToHom import Mathlib.CategoryTheory.Products.Basic /-! diff --git a/Mathlib/CategoryTheory/Monoidal/Category.lean b/Mathlib/CategoryTheory/Monoidal/Category.lean index c38de64d692f4..d80764aa71802 100644 --- a/Mathlib/CategoryTheory/Monoidal/Category.lean +++ b/Mathlib/CategoryTheory/Monoidal/Category.lean @@ -3,6 +3,7 @@ Copyright (c) 2018 Michael Jendrusch. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Michael Jendrusch, Kim Morrison, Bhavik Mehta, Jakob von Raumer -/ +import Mathlib.CategoryTheory.EqToHom import Mathlib.CategoryTheory.Functor.Trifunctor import Mathlib.CategoryTheory.Products.Basic diff --git a/Mathlib/CategoryTheory/Products/Basic.lean b/Mathlib/CategoryTheory/Products/Basic.lean index c38b430629b3b..58958ff88a2de 100644 --- a/Mathlib/CategoryTheory/Products/Basic.lean +++ b/Mathlib/CategoryTheory/Products/Basic.lean @@ -3,7 +3,6 @@ Copyright (c) 2017 Kim Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Stephen Morgan, Kim Morrison -/ -import Mathlib.CategoryTheory.EqToHom import Mathlib.CategoryTheory.Functor.Const import Mathlib.CategoryTheory.Opposites import Mathlib.Data.Prod.Basic @@ -296,7 +295,18 @@ end Equivalence /-- `F.flip` composed with evaluation is the same as evaluating `F`. -/ @[simps!] def flipCompEvaluation (F : A ⥤ B ⥤ C) (a) : F.flip ⋙ (evaluation _ _).obj a ≅ F.obj a := - NatIso.ofComponents fun b => eqToIso rfl + NatIso.ofComponents fun b => Iso.refl _ + +theorem flip_comp_evaluation (F : A ⥤ B ⥤ C) (a) : F.flip ⋙ (evaluation _ _).obj a = F.obj a := + rfl + +/-- `F` composed with evaluation is the same as evaluating `F.flip`. -/ +@[simps!] +def compEvaluation (F : A ⥤ B ⥤ C) (b) : F ⋙ (evaluation _ _).obj b ≅ F.flip.obj b := + NatIso.ofComponents fun a => Iso.refl _ + +theorem comp_evaluation (F : A ⥤ B ⥤ C) (b) : F ⋙ (evaluation _ _).obj b = F.flip.obj b := + rfl variable (A B C) From cbe02ad0a6243d7688e60d69fd7ee0387d6f8059 Mon Sep 17 00:00:00 2001 From: kkytola Date: Mon, 7 Oct 2024 09:08:05 +0000 Subject: [PATCH 08/43] feat: in separable Borel spaces, convergence in distribution is metrizable (not just pseudometrizable) (#15478) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit In earlier PRs it was shown that the Lévy-Prokhorov distance pseudometrizes the convergence in distribution (i.e. weak convergence of probability measures) on separable pseudometric spaces. This PR strengthens the conclusion: the LP distance on probability measures on Borel spaces is a metric (not just a pseudometric), and therefore in separable Borel spaces, convergence in distribution is metrizable (not just pseudometrizable). Co-authored-by: kkytola <39528102+kkytola@users.noreply.github.com> --- .../Measure/LevyProkhorovMetric.lean | 176 +++++++++++------- 1 file changed, 112 insertions(+), 64 deletions(-) diff --git a/Mathlib/MeasureTheory/Measure/LevyProkhorovMetric.lean b/Mathlib/MeasureTheory/Measure/LevyProkhorovMetric.lean index 5b7eb39f8469e..fd684e18cafc6 100644 --- a/Mathlib/MeasureTheory/Measure/LevyProkhorovMetric.lean +++ b/Mathlib/MeasureTheory/Measure/LevyProkhorovMetric.lean @@ -28,10 +28,6 @@ import Mathlib.MeasureTheory.Integral.BoundedContinuousFunction probability measures on a separable space coincides with the topology of convergence in distribution, and in particular convergence in distribution is then pseudometrizable. -## TODO - -* Show that in Borel spaces, the Lévy-Prokhorov distance is a metric; not just a pseudometric. - ## Tags finite measure, probability measure, weak convergence, convergence in distribution, metrizability @@ -198,6 +194,9 @@ lemma levyProkhorovDist_triangle [OpensMeasurableSpace Ω] (μ ν κ : Measure when they are to be equipped with the Lévy-Prokhorov distance. -/ def LevyProkhorov (α : Type*) := α +/-- The "identity" equivalence between the type synonym `LevyProkhorov α` and `α`. -/ +def LevyProkhorov.equiv (α : Type*) : LevyProkhorov α ≃ α := Equiv.refl _ + variable [OpensMeasurableSpace Ω] /-- The Lévy-Prokhorov distance `levyProkhorovEDist` makes `Measure Ω` a pseudoemetric @@ -219,9 +218,43 @@ noncomputable instance levyProkhorovDist_pseudoMetricSpace_finiteMeasure : dist_triangle μ ν κ := levyProkhorovDist_triangle _ _ _ edist_dist μ ν := by simp [← ENNReal.ofReal_coe_nnreal] -/-- The Lévy-Prokhorov distance `levyProkhorovDist` makes `ProbabilityMeasure Ω` a pseudoemetric +lemma measure_le_measure_closure_of_levyProkhorovEDist_eq_zero {μ ν : Measure Ω} + (hLP : levyProkhorovEDist μ ν = 0) {s : Set Ω} (s_mble : MeasurableSet s) + (h_finite : ∃ δ > 0, ν (thickening δ s) ≠ ∞) : + μ s ≤ ν (closure s) := by + have key : Tendsto (fun ε ↦ ν (thickening ε.toReal s)) (𝓝[>] (0 : ℝ≥0∞)) (𝓝 (ν (closure s))) := by + have aux : Tendsto ENNReal.toReal (𝓝[>] 0) (𝓝[>] 0) := by + apply tendsto_nhdsWithin_of_tendsto_nhds_of_eventually_within (s := Ioi 0) ENNReal.toReal + · exact tendsto_nhdsWithin_of_tendsto_nhds (continuousAt_toReal zero_ne_top).tendsto + · filter_upwards [Ioo_mem_nhdsWithin_Ioi ⟨le_rfl, zero_lt_one⟩] with x hx + exact toReal_pos hx.1.ne.symm <| ne_top_of_lt hx.2 + exact (tendsto_measure_thickening h_finite).comp aux + have obs := Tendsto.add key (tendsto_nhdsWithin_of_tendsto_nhds tendsto_id) + simp only [id_eq, add_zero] at obs + apply ge_of_tendsto (b := μ s) obs + filter_upwards [self_mem_nhdsWithin] with ε ε_pos + exact left_measure_le_of_levyProkhorovEDist_lt (B_mble := s_mble) (hLP ▸ ε_pos) + +/-- Two measures at vanishing Lévy-Prokhorov distance from each other assign the same values to all +closed sets. -/ +lemma measure_eq_measure_of_levyProkhorovEDist_eq_zero_of_isClosed {μ ν : Measure Ω} + (hLP : levyProkhorovEDist μ ν = 0) {s : Set Ω} (s_closed : IsClosed s) + (hμs : ∃ δ > 0, μ (thickening δ s) ≠ ∞) (hνs : ∃ δ > 0, ν (thickening δ s) ≠ ∞) : + μ s = ν s := by + apply le_antisymm + · exact measure_le_measure_closure_of_levyProkhorovEDist_eq_zero + hLP s_closed.measurableSet hνs |>.trans <| + le_of_eq (congr_arg _ s_closed.closure_eq) + · exact measure_le_measure_closure_of_levyProkhorovEDist_eq_zero + (levyProkhorovEDist_comm μ ν ▸ hLP) s_closed.measurableSet hμs |>.trans <| + le_of_eq (congr_arg _ s_closed.closure_eq) + +/-- The Lévy-Prokhorov distance `levyProkhorovDist` makes `ProbabilityMeasure Ω` a pseudometric space. The instance is recorded on the type synonym -`LevyProkhorov (ProbabilityMeasure Ω) := ProbabilityMeasure Ω`. -/ +`LevyProkhorov (ProbabilityMeasure Ω) := ProbabilityMeasure Ω`. + +Note: For this pseudometric to give the topology of convergence in distribution, one must +furthermore assume that `Ω` is separable. -/ noncomputable instance levyProkhorovDist_pseudoMetricSpace_probabilityMeasure : PseudoMetricSpace (LevyProkhorov (ProbabilityMeasure Ω)) where dist μ ν := levyProkhorovDist μ.toMeasure ν.toMeasure @@ -233,6 +266,29 @@ noncomputable instance levyProkhorovDist_pseudoMetricSpace_probabilityMeasure : lemma LevyProkhorov.dist_def (μ ν : LevyProkhorov (ProbabilityMeasure Ω)) : dist μ ν = levyProkhorovDist μ.toMeasure ν.toMeasure := rfl +/-- If `Ω` is a Borel space, then the Lévy-Prokhorov distance `levyProkhorovDist` makes +`ProbabilityMeasure Ω` a metric space. The instance is recorded on the type synonym +`LevyProkhorov (ProbabilityMeasure Ω) := ProbabilityMeasure Ω`. + +Note: For this metric to give the topology of convergence in distribution, one must +furthermore assume that `Ω` is separable. -/ +noncomputable instance levyProkhorovDist_metricSpace_probabilityMeasure [BorelSpace Ω] : + MetricSpace (LevyProkhorov (ProbabilityMeasure Ω)) where + eq_of_dist_eq_zero := by + intro μ ν h + apply (LevyProkhorov.equiv _).injective + apply ProbabilityMeasure.toMeasure_injective + apply ext_of_generate_finite _ ?_ isPiSystem_isClosed ?_ (by simp) + · rw [BorelSpace.measurable_eq (α := Ω), borel_eq_generateFrom_isClosed] + · intro A A_closed + apply measure_eq_measure_of_levyProkhorovEDist_eq_zero_of_isClosed + · simpa only [levyProkhorovEDist_ne_top μ.toMeasure ν.toMeasure, mem_setOf_eq, + or_false, ne_eq, zero_ne_top, not_false_eq_true, zero_toReal] + using (toReal_eq_zero_iff _).mp h + · exact A_closed + · exact ⟨1, Real.zero_lt_one, measure_ne_top _ _⟩ + · exact ⟨1, Real.zero_lt_one, measure_ne_top _ _⟩ + /-- A simple sufficient condition for bounding `levyProkhorovEDist` between probability measures from above. The condition involves only one of two natural bounds, the other bound is for free. -/ lemma levyProkhorovEDist_le_of_forall_le @@ -279,20 +335,6 @@ open BoundedContinuousFunction variable {ι : Type*} {Ω : Type*} [MeasurableSpace Ω] -/-- Coercion from the type synonym `LevyProkhorov (ProbabilityMeasure Ω)` -to `ProbabilityMeasure Ω`. -/ -def LevyProkhorov.toProbabilityMeasure (μ : LevyProkhorov (ProbabilityMeasure Ω)) : - ProbabilityMeasure Ω := μ - -/-- Coercion to the type synonym `LevyProkhorov (ProbabilityMeasure Ω)` -from `ProbabilityMeasure Ω`. -/ -def ProbabilityMeasure.toLevyProkhorov (μ : ProbabilityMeasure Ω) : - LevyProkhorov (ProbabilityMeasure Ω) := μ - -/-- Coercion from the type synonym `LevyProkhorov (FiniteMeasure Ω)` to `FiniteMeasure Ω`. -/ -def LevyProkhorov.finiteMeasure (μ : LevyProkhorov (FiniteMeasure Ω)) : - FiniteMeasure Ω := μ - variable [PseudoMetricSpace Ω] [OpensMeasurableSpace Ω] /-- A version of the layer cake formula for bounded continuous functions which have finite integral: @@ -384,13 +426,13 @@ lemma tendsto_integral_meas_thickening_le (f : Ω →ᵇ ℝ) · exact isClosed_le continuous_const f.continuous · exact measure_ne_top _ _ -/-- The coercion `LevyProkhorov (ProbabilityMeasure Ω) → ProbabilityMeasure Ω` is continuous. -/ -lemma LevyProkhorov.continuous_toProbabilityMeasure : - Continuous (LevyProkhorov.toProbabilityMeasure (Ω := Ω)) := by +/-- The identity map `LevyProkhorov (ProbabilityMeasure Ω) → ProbabilityMeasure Ω` is continuous. -/ +lemma LevyProkhorov.continuous_equiv_probabilityMeasure : + Continuous (LevyProkhorov.equiv (α := ProbabilityMeasure Ω)) := by refine SeqContinuous.continuous ?_ intro μs ν hμs - set P := ν.toProbabilityMeasure -- more palatable notation - set Ps := fun n ↦ (μs n).toProbabilityMeasure -- more palatable notation + set P := LevyProkhorov.equiv _ ν -- more palatable notation + set Ps := fun n ↦ LevyProkhorov.equiv _ (μs n) -- more palatable notation rw [ProbabilityMeasure.tendsto_iff_forall_integral_tendsto] refine fun f ↦ tendsto_integral_of_forall_limsup_integral_le_integral ?_ f intro f f_nn @@ -433,9 +475,8 @@ lemma LevyProkhorov.continuous_toProbabilityMeasure : · rw [ENNReal.ofReal_add (by positivity) (by positivity), ← add_zero (levyProkhorovEDist _ _)] apply ENNReal.add_lt_add_of_le_of_lt (levyProkhorovEDist_ne_top _ _) (le_of_eq ?_) (ofReal_pos.mpr εs_pos) - rw [LevyProkhorov.dist_def, levyProkhorovDist, - ofReal_toReal (levyProkhorovEDist_ne_top _ _)] - simp only [Ps, P, LevyProkhorov.toProbabilityMeasure] + rw [LevyProkhorov.dist_def, levyProkhorovDist, ofReal_toReal (levyProkhorovEDist_ne_top _ _)] + rfl · exact Eventually.of_forall f_nn · simp only [IsCoboundedUnder, IsCobounded, eventually_map, eventually_atTop, forall_exists_index] @@ -444,9 +485,9 @@ lemma LevyProkhorov.continuous_toProbabilityMeasure : /-- The topology of the Lévy-Prokhorov metric is at least as fine as the topology of convergence in distribution. -/ theorem levyProkhorov_le_convergenceInDistribution : - TopologicalSpace.coinduced (LevyProkhorov.toProbabilityMeasure (Ω := Ω)) inferInstance + TopologicalSpace.coinduced (LevyProkhorov.equiv (α := ProbabilityMeasure Ω)) inferInstance ≤ (inferInstance : TopologicalSpace (ProbabilityMeasure Ω)) := - (LevyProkhorov.continuous_toProbabilityMeasure).coinduced_le + (LevyProkhorov.continuous_equiv_probabilityMeasure).coinduced_le end Levy_Prokhorov_is_finer @@ -456,13 +497,34 @@ section Levy_Prokhorov_metrizes_convergence_in_distribution open BoundedContinuousFunction TopologicalSpace -variable {ι : Type*} (Ω : Type*) [PseudoMetricSpace Ω] +variable {ι : Type*} {Ω : Type*} [PseudoMetricSpace Ω] variable [MeasurableSpace Ω] [OpensMeasurableSpace Ω] +lemma ProbabilityMeasure.toMeasure_add_pos_gt_mem_nhds (P : ProbabilityMeasure Ω) + {G : Set Ω} (G_open : IsOpen G) {ε : ℝ≥0∞} (ε_pos : 0 < ε) : + {Q | P.toMeasure G < Q.toMeasure G + ε} ∈ 𝓝 P := by + by_cases easy : P.toMeasure G < ε + · exact Eventually.of_forall (fun _ ↦ lt_of_lt_of_le easy le_add_self) + by_cases ε_top : ε = ∞ + · simp [ε_top, measure_lt_top] + simp only [not_lt] at easy + have aux : P.toMeasure G - ε < liminf (fun Q ↦ Q.toMeasure G) (𝓝 P) := by + apply lt_of_lt_of_le (ENNReal.sub_lt_self (measure_lt_top _ _).ne _ _) + <| ProbabilityMeasure.le_liminf_measure_open_of_tendsto tendsto_id G_open + · exact (lt_of_lt_of_le ε_pos easy).ne.symm + · exact ε_pos.ne.symm + filter_upwards [gt_mem_sets_of_limsInf_gt (α := ℝ≥0∞) isBounded_ge_of_bot + (show P.toMeasure G - ε < limsInf ((𝓝 P).map (fun Q ↦ Q.toMeasure G)) from aux)] with Q hQ + simp only [preimage_setOf_eq, mem_setOf_eq] at hQ + convert ENNReal.add_lt_add_right ε_top hQ + exact (tsub_add_cancel_of_le easy).symm + +variable [SeparableSpace Ω] + +variable (Ω) in /-- In a separable pseudometric space, for any ε > 0 there exists a countable collection of disjoint Borel measurable subsets of diameter at most ε that cover the whole space. -/ -lemma SeparableSpace.exists_measurable_partition_diam_le [SeparableSpace Ω] - {ε : ℝ} (ε_pos : 0 < ε) : +lemma SeparableSpace.exists_measurable_partition_diam_le {ε : ℝ} (ε_pos : 0 < ε) : ∃ (As : ℕ → Set Ω), (∀ n, MeasurableSet (As n)) ∧ (∀ n, Bornology.IsBounded (As n)) ∧ (∀ n, diam (As n) ≤ ε) ∧ (⋃ n, As n = univ) ∧ (Pairwise (fun (n m : ℕ) ↦ Disjoint (As n) (As m))) := by @@ -492,29 +554,8 @@ lemma SeparableSpace.exists_measurable_partition_diam_le [SeparableSpace Ω] simpa only [← aux] using iUnion_disjointed · exact disjoint_disjointed Bs -variable {Ω} - -lemma ProbabilityMeasure.toMeasure_add_pos_gt_mem_nhds (P : ProbabilityMeasure Ω) - {G : Set Ω} (G_open : IsOpen G) {ε : ℝ≥0∞} (ε_pos : 0 < ε) : - {Q | P.toMeasure G < Q.toMeasure G + ε} ∈ 𝓝 P := by - by_cases easy : P.toMeasure G < ε - · exact Eventually.of_forall (fun _ ↦ lt_of_lt_of_le easy le_add_self) - by_cases ε_top : ε = ∞ - · simp [ε_top, measure_lt_top] - simp only [not_lt] at easy - have aux : P.toMeasure G - ε < liminf (fun Q ↦ Q.toMeasure G) (𝓝 P) := by - apply lt_of_lt_of_le (ENNReal.sub_lt_self (measure_lt_top _ _).ne _ _) - <| ProbabilityMeasure.le_liminf_measure_open_of_tendsto tendsto_id G_open - · exact (lt_of_lt_of_le ε_pos easy).ne.symm - · exact ε_pos.ne.symm - filter_upwards [gt_mem_sets_of_limsInf_gt (α := ℝ≥0∞) isBounded_ge_of_bot - (show P.toMeasure G - ε < limsInf ((𝓝 P).map (fun Q ↦ Q.toMeasure G)) from aux)] with Q hQ - simp only [preimage_setOf_eq, mem_setOf_eq] at hQ - convert ENNReal.add_lt_add_right ε_top hQ - exact (tsub_add_cancel_of_le easy).symm - -lemma ProbabilityMeasure.continuous_toLevyProkhorov [SeparableSpace Ω] : - Continuous (ProbabilityMeasure.toLevyProkhorov (Ω := Ω)) := by +lemma LevyProkhorov.continuous_equiv_symm_probabilityMeasure : + Continuous (LevyProkhorov.equiv (α := ProbabilityMeasure Ω)).symm := by -- We check continuity of `id : ProbabilityMeasure Ω → LevyProkhorov (ProbabilityMeasure Ω)` at -- each point `P : ProbabilityMeasure Ω`. rw [continuous_iff_continuousAt] @@ -612,22 +653,22 @@ lemma ProbabilityMeasure.continuous_toLevyProkhorov [SeparableSpace Ω] : /-- The topology of the Lévy-Prokhorov metric on probability measures on a separable space coincides with the topology of convergence in distribution. -/ -theorem levyProkhorov_eq_convergenceInDistribution [SeparableSpace Ω] : +theorem levyProkhorov_eq_convergenceInDistribution : (inferInstance : TopologicalSpace (ProbabilityMeasure Ω)) - = TopologicalSpace.coinduced (LevyProkhorov.toProbabilityMeasure (Ω := Ω)) inferInstance := - le_antisymm (ProbabilityMeasure.continuous_toLevyProkhorov (Ω := Ω)).coinduced_le + = TopologicalSpace.coinduced (LevyProkhorov.equiv _) inferInstance := + le_antisymm (LevyProkhorov.continuous_equiv_symm_probabilityMeasure (Ω := Ω)).coinduced_le levyProkhorov_le_convergenceInDistribution /-- The identity map is a homeomorphism from `ProbabilityMeasure Ω` with the topology of convergence in distribution to `ProbabilityMeasure Ω` with the Lévy-Prokhorov (pseudo)metric. -/ -def homeomorph_probabilityMeasure_levyProkhorov [SeparableSpace Ω] : +def homeomorph_probabilityMeasure_levyProkhorov : ProbabilityMeasure Ω ≃ₜ LevyProkhorov (ProbabilityMeasure Ω) where - toFun := ProbabilityMeasure.toLevyProkhorov (Ω := Ω) - invFun := LevyProkhorov.toProbabilityMeasure (Ω := Ω) + toFun := LevyProkhorov.equiv _ + invFun := (LevyProkhorov.equiv _).symm left_inv := congrFun rfl right_inv := congrFun rfl - continuous_toFun := ProbabilityMeasure.continuous_toLevyProkhorov - continuous_invFun := LevyProkhorov.continuous_toProbabilityMeasure + continuous_toFun := LevyProkhorov.continuous_equiv_symm_probabilityMeasure + continuous_invFun := LevyProkhorov.continuous_equiv_probabilityMeasure /-- The topology of convergence in distribution on a separable space is pseudo-metrizable. -/ instance (X : Type*) [TopologicalSpace X] [PseudoMetrizableSpace X] [SeparableSpace X] @@ -636,6 +677,13 @@ instance (X : Type*) [TopologicalSpace X] [PseudoMetrizableSpace X] [SeparableSp letI : PseudoMetricSpace X := TopologicalSpace.pseudoMetrizableSpacePseudoMetric X (homeomorph_probabilityMeasure_levyProkhorov (Ω := X)).inducing.pseudoMetrizableSpace +/-- The topology of convergence in distribution on a separable Borel space is metrizable. -/ +instance instMetrizableSpaceProbabilityMeasure (X : Type*) [TopologicalSpace X] + [PseudoMetrizableSpace X] [SeparableSpace X] [MeasurableSpace X] [BorelSpace X] : + MetrizableSpace (ProbabilityMeasure X) := by + letI : PseudoMetricSpace X := TopologicalSpace.pseudoMetrizableSpacePseudoMetric X + exact homeomorph_probabilityMeasure_levyProkhorov.embedding.metrizableSpace + end Levy_Prokhorov_metrizes_convergence_in_distribution end MeasureTheory -- namespace From b18f6b58af9dba6bbd877629cf7c6809c049ff1e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Mon, 7 Oct 2024 10:11:57 +0000 Subject: [PATCH 09/43] feat(Data/Finsupp/Defs): `induction_on_max` (#16061) A finitely supported function can be built in ascending or descending order of "degree". --- Mathlib/Data/Finsupp/Defs.lean | 54 ++++++++++++++++++++++++++++++++++ 1 file changed, 54 insertions(+) diff --git a/Mathlib/Data/Finsupp/Defs.lean b/Mathlib/Data/Finsupp/Defs.lean index 70c91059b3738..11717314f657f 100644 --- a/Mathlib/Data/Finsupp/Defs.lean +++ b/Mathlib/Data/Finsupp/Defs.lean @@ -1048,6 +1048,60 @@ theorem induction_linear {p : (α →₀ M) → Prop} (f : α →₀ M) (h0 : p (hadd : ∀ f g : α →₀ M, p f → p g → p (f + g)) (hsingle : ∀ a b, p (single a b)) : p f := induction₂ f h0 fun _a _b _f _ _ w => hadd _ _ w (hsingle _ _) +section LinearOrder + +variable [LinearOrder α] {p : (α →₀ M) → Prop} + +/-- A finitely supported function can be built by adding up `single a b` for increasing `a`. + +The theorem `induction_on_max₂` swaps the argument order in the sum. -/ +theorem induction_on_max (f : α →₀ M) (h0 : p 0) + (ha : ∀ (a b) (f : α →₀ M), (∀ c ∈ f.support, c < a) → b ≠ 0 → p f → p (single a b + f)) : + p f := by + suffices ∀ (s) (f : α →₀ M), f.support = s → p f from this _ _ rfl + refine fun s => s.induction_on_max (fun f h => ?_) (fun a s hm hf f hs => ?_) + · rwa [support_eq_empty.1 h] + · have hs' : (erase a f).support = s := by + rw [support_erase, hs, erase_insert (fun ha => (hm a ha).false)] + rw [← single_add_erase a f] + refine ha _ _ _ (fun c hc => hm _ <| hs'.symm ▸ hc) ?_ (hf _ hs') + rw [← mem_support_iff, hs] + exact mem_insert_self a s + +/-- A finitely supported function can be built by adding up `single a b` for decreasing `a`. + +The theorem `induction_on_min₂` swaps the argument order in the sum. -/ +theorem induction_on_min (f : α →₀ M) (h0 : p 0) + (ha : ∀ (a b) (f : α →₀ M), (∀ c ∈ f.support, a < c) → b ≠ 0 → p f → p (single a b + f)) : + p f := + induction_on_max (α := αᵒᵈ) f h0 ha + +/-- A finitely supported function can be built by adding up `single a b` for increasing `a`. + +The theorem `induction_on_max` swaps the argument order in the sum. -/ +theorem induction_on_max₂ (f : α →₀ M) (h0 : p 0) + (ha : ∀ (a b) (f : α →₀ M), (∀ c ∈ f.support, c < a) → b ≠ 0 → p f → p (f + single a b)) : + p f := by + suffices ∀ (s) (f : α →₀ M), f.support = s → p f from this _ _ rfl + refine fun s => s.induction_on_max (fun f h => ?_) (fun a s hm hf f hs => ?_) + · rwa [support_eq_empty.1 h] + · have hs' : (erase a f).support = s := by + rw [support_erase, hs, erase_insert (fun ha => (hm a ha).false)] + rw [← erase_add_single a f] + refine ha _ _ _ (fun c hc => hm _ <| hs'.symm ▸ hc) ?_ (hf _ hs') + rw [← mem_support_iff, hs] + exact mem_insert_self a s + +/-- A finitely supported function can be built by adding up `single a b` for decreasing `a`. + +The theorem `induction_on_min` swaps the argument order in the sum. -/ +theorem induction_on_min₂ (f : α →₀ M) (h0 : p 0) + (ha : ∀ (a b) (f : α →₀ M), (∀ c ∈ f.support, a < c) → b ≠ 0 → p f → p (f + single a b)) : + p f := + induction_on_max₂ (α := αᵒᵈ) f h0 ha + +end LinearOrder + @[simp] theorem add_closure_setOf_eq_single : AddSubmonoid.closure { f : α →₀ M | ∃ a b, f = single a b } = ⊤ := From 2eaa58a74f623e9e02f387eaf20a7a57a5e00721 Mon Sep 17 00:00:00 2001 From: judithludwig Date: Mon, 7 Oct 2024 10:11:58 +0000 Subject: [PATCH 10/43] feat(LinearAlgebra/Matrix/GeneralLinearGroup): add lemmas on hom induced from a ring hom (#17433) This PR adds some basic lemmas on the hom between GL(n)s induced from a ring hom to LinearAlgebra.Matrix.GeneralLinearGroup.Defs.lean . Co-authored-by: Christian Merten (github: chrisflav). Used in project Formalization of the Bruhat-Tits Tree. Co-authored-by: judithludwig <150042763+judithludwig@users.noreply.github.com> --- .../Matrix/GeneralLinearGroup/Defs.lean | 42 +++++++++++++++++++ 1 file changed, 42 insertions(+) diff --git a/Mathlib/LinearAlgebra/Matrix/GeneralLinearGroup/Defs.lean b/Mathlib/LinearAlgebra/Matrix/GeneralLinearGroup/Defs.lean index 50c1c2dfc8bb0..5ae49bf1b77c9 100644 --- a/Mathlib/LinearAlgebra/Matrix/GeneralLinearGroup/Defs.lean +++ b/Mathlib/LinearAlgebra/Matrix/GeneralLinearGroup/Defs.lean @@ -137,6 +137,10 @@ def map (f : R →+* S) : GL n R →* GL n S := Units.map <| (RingHom.mapMatrix theorem map_id : map (RingHom.id R) = MonoidHom.id (GL n R) := rfl +@[simp] +protected lemma map_apply (f : R →+* S) (i j : n) (g : GL n R) : map f g i j = f (g i j) := by + rfl + @[simp] theorem map_comp (f : T →+* R) (g : R →+* S) : map (g.comp f) = (map g).comp (map (n := n) f) := @@ -147,6 +151,44 @@ theorem map_comp_apply (f : T →+* R) (g : R →+* S) (x : GL n T) : (map g).comp (map f) x = map g (map f x) := rfl +variable (f : R →+* S) + +@[simp] +protected lemma map_one : map f (1 : GL n R) = 1 := by + ext + simp only [_root_.map_one, Units.val_one] + +protected lemma map_mul (g h : GL n R) : map f (g * h) = map f g * map f h := by + ext + simp only [_root_.map_mul, Units.val_mul] + +protected lemma map_inv (g : GL n R) : map f g⁻¹ = (map f g)⁻¹ := by + ext + simp only [_root_.map_inv, coe_units_inv] + +protected lemma map_det (g : GL n R) : Matrix.GeneralLinearGroup.det (map f g) = + Units.map f (Matrix.GeneralLinearGroup.det g) := by + ext + simp only [map, RingHom.mapMatrix_apply, Units.inv_eq_val_inv, Matrix.coe_units_inv, + Matrix.GeneralLinearGroup.val_det_apply, Units.coe_map, MonoidHom.coe_coe] + exact Eq.symm (RingHom.map_det f g.1) + +lemma map_mul_map_inv (g : GL n R) : map f g * map f g⁻¹ = 1 := by + simp only [map_inv, mul_inv_cancel] + +lemma map_inv_mul_map (g : GL n R) : map f g⁻¹ * map f g = 1 := by + simp only [map_inv, inv_mul_cancel] + +@[simp] +lemma coe_map_mul_map_inv (g : GL n R) : g.val.map f * g.val⁻¹.map f = 1 := by + rw [← Matrix.map_mul] + simp only [isUnits_det_units, mul_nonsing_inv, map_zero, _root_.map_one, Matrix.map_one] + +@[simp] +lemma coe_map_inv_mul_map (g : GL n R) : g.val⁻¹.map f * g.val.map f = 1 := by + rw [← Matrix.map_mul] + simp only [isUnits_det_units, nonsing_inv_mul, map_zero, _root_.map_one, Matrix.map_one] + end GeneralLinearGroup namespace SpecialLinearGroup From 140847a5ec73c04c49a8b93b632675ee2b9e9700 Mon Sep 17 00:00:00 2001 From: damiano Date: Mon, 7 Oct 2024 11:34:31 +0000 Subject: [PATCH 11/43] CI: temporary fix to CI not finding cache (#17490) --- .github/build.in.yml | 2 +- .github/workflows/bors.yml | 2 +- .github/workflows/build.yml | 2 +- .github/workflows/build_fork.yml | 2 +- 4 files changed, 4 insertions(+), 4 deletions(-) diff --git a/.github/build.in.yml b/.github/build.in.yml index c3b0ebc6aad3d..1b0a8caaa4199 100644 --- a/.github/build.in.yml +++ b/.github/build.in.yml @@ -133,7 +133,7 @@ jobs: run: | rm -rf .lake/build/lib/Mathlib/ # Fail quickly if the cache is completely cold, by checking for Mathlib.Init - lake exe cache get Mathlib.Init + lake exe cache get #Mathlib.Init lake build --no-build Mathlib.Init && lake exe cache get || echo "No cache for 'Mathlib.Init' available" - name: update {Mathlib, Tactic, Counterexamples, Archive}.lean diff --git a/.github/workflows/bors.yml b/.github/workflows/bors.yml index e8f0399a83c22..82c9c075d5fd7 100644 --- a/.github/workflows/bors.yml +++ b/.github/workflows/bors.yml @@ -143,7 +143,7 @@ jobs: run: | rm -rf .lake/build/lib/Mathlib/ # Fail quickly if the cache is completely cold, by checking for Mathlib.Init - lake exe cache get Mathlib.Init + lake exe cache get #Mathlib.Init lake build --no-build Mathlib.Init && lake exe cache get || echo "No cache for 'Mathlib.Init' available" - name: update {Mathlib, Tactic, Counterexamples, Archive}.lean diff --git a/.github/workflows/build.yml b/.github/workflows/build.yml index 72978c1079fe8..5ae06124e19ce 100644 --- a/.github/workflows/build.yml +++ b/.github/workflows/build.yml @@ -150,7 +150,7 @@ jobs: run: | rm -rf .lake/build/lib/Mathlib/ # Fail quickly if the cache is completely cold, by checking for Mathlib.Init - lake exe cache get Mathlib.Init + lake exe cache get #Mathlib.Init lake build --no-build Mathlib.Init && lake exe cache get || echo "No cache for 'Mathlib.Init' available" - name: update {Mathlib, Tactic, Counterexamples, Archive}.lean diff --git a/.github/workflows/build_fork.yml b/.github/workflows/build_fork.yml index 3bfe7e18fee50..e3e86f31e2e1e 100644 --- a/.github/workflows/build_fork.yml +++ b/.github/workflows/build_fork.yml @@ -147,7 +147,7 @@ jobs: run: | rm -rf .lake/build/lib/Mathlib/ # Fail quickly if the cache is completely cold, by checking for Mathlib.Init - lake exe cache get Mathlib.Init + lake exe cache get #Mathlib.Init lake build --no-build Mathlib.Init && lake exe cache get || echo "No cache for 'Mathlib.Init' available" - name: update {Mathlib, Tactic, Counterexamples, Archive}.lean From fcafb6da95654f09ef5dde400321f97c5b57f69a Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Mon, 7 Oct 2024 11:43:29 +0000 Subject: [PATCH 12/43] feat(Data/List/Chain): lemmas on `Chain'` and `replicate` (#17106) Co-authored-by: Bhavik Mehta --- Mathlib/Data/List/Chain.lean | 24 ++++++++++++++++++++++++ 1 file changed, 24 insertions(+) diff --git a/Mathlib/Data/List/Chain.lean b/Mathlib/Data/List/Chain.lean index 1213160f03d0a..a627cab57f2a6 100644 --- a/Mathlib/Data/List/Chain.lean +++ b/Mathlib/Data/List/Chain.lean @@ -142,6 +142,19 @@ theorem chain_iff_get {R} : ∀ {a : α} {l : List α}, Chain R a l ↔ intro i w exact h (i+1) (by simp only [length_cons]; omega) +theorem chain_replicate_of_rel (n : ℕ) {a : α} (h : r a a) : Chain r a (replicate n a) := + match n with + | 0 => Chain.nil + | n + 1 => Chain.cons h (chain_replicate_of_rel n h) + +theorem chain_eq_iff_eq_replicate {a : α} {l : List α} : + Chain (· = ·) a l ↔ l = replicate l.length a := + match l with + | [] => by simp + | b :: l => by + rw [chain_cons] + simp (config := {contextual := true}) [eq_comm, replicate_succ, chain_eq_iff_eq_replicate] + theorem Chain'.imp {S : α → α → Prop} (H : ∀ a b, R a b → S a b) {l : List α} (p : Chain' R l) : Chain' S l := by cases l <;> [trivial; exact Chain.imp H p] @@ -435,6 +448,17 @@ lemma Chain'.iterate_eq_of_apply_eq {α : Type*} {f : α → α} {l : List α} apply hl omega +theorem chain'_replicate_of_rel (n : ℕ) {a : α} (h : r a a) : Chain' r (replicate n a) := + match n with + | 0 => chain'_nil + | n + 1 => chain_replicate_of_rel n h + +theorem chain'_eq_iff_eq_replicate {l : List α} : + Chain' (· = ·) l ↔ ∀ a ∈ l.head?, l = replicate l.length a := + match l with + | [] => by simp + | a :: l => by simp [Chain', chain_eq_iff_eq_replicate, replicate_succ] + end List From c8a595b92cfc5144c5a75d8a158f57b8c711322b Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 7 Oct 2024 11:43:30 +0000 Subject: [PATCH 13/43] fix: add back `List.decidableInfix` (#17420) This was deleted [in the bump to Lean v4.12.0](https://github.com/leanprover-community/mathlib4/pull/16433/files#r1787921652) in #16433, but was not actually upstreamed. I tweaked the proof slightly to not mention the instance names. --- Mathlib/Data/List/Infix.lean | 8 ++++++++ 1 file changed, 8 insertions(+) diff --git a/Mathlib/Data/List/Infix.lean b/Mathlib/Data/List/Infix.lean index 7e655efa8abee..a399857f66da2 100644 --- a/Mathlib/Data/List/Infix.lean +++ b/Mathlib/Data/List/Infix.lean @@ -78,6 +78,14 @@ theorem concat_get_prefix {x y : List α} (h : x <+: y) (hl : x.length < y.lengt convert List.take_append_drop (x.length + 1) y using 2 rw [← List.take_concat_get, List.concat_eq_append]; rfl +instance decidableInfix [DecidableEq α] : ∀ l₁ l₂ : List α, Decidable (l₁ <:+: l₂) + | [], l₂ => isTrue ⟨[], l₂, rfl⟩ + | a :: l₁, [] => isFalse fun ⟨s, t, te⟩ => by simp at te + | l₁, b :: l₂ => + letI := l₁.decidableInfix l₂ + @decidable_of_decidable_of_iff (l₁ <+: b :: l₂ ∨ l₁ <:+: l₂) _ _ + infix_cons_iff.symm + @[deprecated cons_prefix_cons (since := "2024-08-14")] theorem cons_prefix_iff : a :: l₁ <+: b :: l₂ ↔ a = b ∧ l₁ <+: l₂ := by simp From bb5f02d3c6437b66618b9f5dacecf12cc5bc710d Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Mon, 7 Oct 2024 11:43:32 +0000 Subject: [PATCH 14/43] =?UTF-8?q?feat:=20`fun=20x=20=E2=86=A6=20(x=20?= =?UTF-8?q?=E2=8A=93=20z,=20x=20=E2=8A=94=20z)`=20is=20StrictMono=20in=20M?= =?UTF-8?q?odularLattice=20(#17457)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit In the proof of `wellFounded_lt_exact_sequence`, it is implicitly shown that `fun x : α ↦ (x ⊓ z, x ⊔ z)` is StrictMono in a modular lattice, with the lexicographic order on `α × α`. Here we show the stronger result with `α × α` equipped with the product order, and golf the proof. If `G` is not an abelian group, then `α = Subgroup G` is not necessarily a modular lattice. However, if `z` is a normal subgroup, the result still holds true. We prove a closely related result `strictMono_comap_prod_map` which replaces `inf` with `comap` and `sup` with `map`, so it concerns a function to `Subgroup z × Subgroup (G ⧸ z)` instead. The same result for submodules immediately implies that Noetherian/Artinian are closed under extensions, which is proven in #17425 using `wellFounded_lt/gt_exact_sequence`. If `z` is not even normal, the function to `Subgroup z × Set (G ⧸ z)` is still StrictMono. --- Mathlib/GroupTheory/Coset/Basic.lean | 15 +++++++++++++++ Mathlib/GroupTheory/QuotientGroup/Basic.lean | 5 +++++ Mathlib/LinearAlgebra/Quotient.lean | 7 ++++++- Mathlib/Order/ModularLattice.lean | 19 ++++++++----------- 4 files changed, 34 insertions(+), 12 deletions(-) diff --git a/Mathlib/GroupTheory/Coset/Basic.lean b/Mathlib/GroupTheory/Coset/Basic.lean index 6143225812daa..75ec037119a2a 100644 --- a/Mathlib/GroupTheory/Coset/Basic.lean +++ b/Mathlib/GroupTheory/Coset/Basic.lean @@ -402,6 +402,21 @@ theorem out_eq' (a : α ⧸ s) : mk a.out' = a := variable (s) +/-- Given a subgroup `s`, the function that sends a subgroup `t` to the pair consisting of +its intersection with `s` and its image in the quotient `α ⧸ s` is strictly monotone, even though +it is not injective in general. -/ +@[to_additive QuotientAddGroup.strictMono_comap_prod_image "Given an additive subgroup `s`, +the function that sends an additive subgroup `t` to the pair consisting of +its intersection with `s` and its image in the quotient `α ⧸ s` +is strictly monotone, even though it is not injective in general."] +theorem strictMono_comap_prod_image : + StrictMono fun t : Subgroup α ↦ (t.comap s.subtype, mk (s := s) '' t) := by + refine fun t₁ t₂ h ↦ ⟨⟨Subgroup.comap_mono h.1, Set.image_mono h.1⟩, + mt (fun ⟨le1, le2⟩ a ha ↦ ?_) h.2⟩ + obtain ⟨a', h', eq⟩ := le2 ⟨_, ha, rfl⟩ + convert ← t₁.mul_mem h' (@le1 ⟨_, QuotientGroup.eq.1 eq⟩ <| t₂.mul_mem (t₂.inv_mem <| h.1 h') ha) + apply mul_inv_cancel_left + /- It can be useful to write `obtain ⟨h, H⟩ := mk_out'_eq_mul ...`, and then `rw [H]` or `simp_rw [H]` or `simp only [H]`. In order for `simp_rw` and `simp only` to work, this lemma is stated in terms of an arbitrary `h : s`, rather than the specific `h = g⁻¹ * (mk g).out'`. -/ diff --git a/Mathlib/GroupTheory/QuotientGroup/Basic.lean b/Mathlib/GroupTheory/QuotientGroup/Basic.lean index 7857eb080a61d..ec446e552f0d8 100644 --- a/Mathlib/GroupTheory/QuotientGroup/Basic.lean +++ b/Mathlib/GroupTheory/QuotientGroup/Basic.lean @@ -167,6 +167,11 @@ theorem mk_prod {G ι : Type*} [CommGroup G] (N : Subgroup G) (s : Finset ι) {f @[to_additive (attr := simp)] lemma map_mk'_self : N.map (mk' N) = ⊥ := by aesop +@[to_additive QuotientAddGroup.strictMono_comap_prod_map] +theorem strictMono_comap_prod_map : + StrictMono fun H : Subgroup G ↦ (H.comap N.subtype, H.map (mk' N)) := + strictMono_comap_prod_image N + /-- A group homomorphism `φ : G →* M` with `N ⊆ ker(φ)` descends (i.e. `lift`s) to a group homomorphism `G/N →* M`. -/ @[to_additive "An `AddGroup` homomorphism `φ : G →+ M` with `N ⊆ ker(φ)` descends (i.e. `lift`s) diff --git a/Mathlib/LinearAlgebra/Quotient.lean b/Mathlib/LinearAlgebra/Quotient.lean index 61d6fd9576a2a..ca7aacc2896a2 100644 --- a/Mathlib/LinearAlgebra/Quotient.lean +++ b/Mathlib/LinearAlgebra/Quotient.lean @@ -282,9 +282,14 @@ def mkQ : M →ₗ[R] M ⧸ p where theorem mkQ_apply (x : M) : p.mkQ x = (Quotient.mk x : M ⧸ p) := rfl -theorem mkQ_surjective (A : Submodule R M) : Function.Surjective A.mkQ := by +theorem mkQ_surjective : Function.Surjective p.mkQ := by rintro ⟨x⟩; exact ⟨x, rfl⟩ +theorem strictMono_comap_prod_map : + StrictMono fun m : Submodule R M ↦ (m.comap p.subtype, m.map p.mkQ) := + fun m₁ m₂ ↦ QuotientAddGroup.strictMono_comap_prod_map + p.toAddSubgroup (a := m₁.toAddSubgroup) (b := m₂.toAddSubgroup) + end variable {R₂ M₂ : Type*} [Ring R₂] [AddCommGroup M₂] [Module R₂ M₂] {τ₁₂ : R →+* R₂} diff --git a/Mathlib/Order/ModularLattice.lean b/Mathlib/Order/ModularLattice.lean index 6dfc3e2db07ba..7b2b7f77ae7f8 100644 --- a/Mathlib/Order/ModularLattice.lean +++ b/Mathlib/Order/ModularLattice.lean @@ -218,6 +218,10 @@ theorem sup_lt_sup_of_lt_of_inf_le_inf (hxy : x < y) (hinf : y ⊓ z ≤ x ⊓ z theorem inf_lt_inf_of_lt_of_sup_le_sup (hxy : x < y) (hinf : y ⊔ z ≤ x ⊔ z) : x ⊓ z < y ⊓ z := sup_lt_sup_of_lt_of_inf_le_inf (α := αᵒᵈ) hxy hinf +theorem strictMono_inf_prod_sup : StrictMono fun x ↦ (x ⊓ z, x ⊔ z) := fun _x _y hxy ↦ + ⟨⟨inf_le_inf_right _ hxy.le, sup_le_sup_right hxy.le _⟩, + fun ⟨inf_le, sup_le⟩ ↦ (sup_lt_sup_of_lt_of_inf_le_inf hxy inf_le).not_le sup_le⟩ + /-- A generalization of the theorem that if `N` is a submodule of `M` and `N` and `M / N` are both Artinian, then `M` is Artinian. -/ theorem wellFounded_lt_exact_sequence {β γ : Type*} [PartialOrder β] [Preorder γ] @@ -225,16 +229,9 @@ theorem wellFounded_lt_exact_sequence {β γ : Type*} [PartialOrder β] [Preorde (f₁ : β → α) (f₂ : α → β) (g₁ : γ → α) (g₂ : α → γ) (gci : GaloisCoinsertion f₁ f₂) (gi : GaloisInsertion g₂ g₁) (hf : ∀ a, f₁ (f₂ a) = a ⊓ K) (hg : ∀ a, g₁ (g₂ a) = a ⊔ K) : WellFoundedLT α := - ⟨Subrelation.wf - (@fun A B hAB => - show Prod.Lex (· < ·) (· < ·) (f₂ A, g₂ A) (f₂ B, g₂ B) by - simp only [Prod.lex_def, lt_iff_le_not_le, ← gci.l_le_l_iff, ← gi.u_le_u_iff, hf, hg, - le_antisymm_iff] - simp only [gci.l_le_l_iff, gi.u_le_u_iff, ← lt_iff_le_not_le, ← le_antisymm_iff] - rcases lt_or_eq_of_le (inf_le_inf_right K (le_of_lt hAB)) with h | h - · exact Or.inl h - · exact Or.inr ⟨h, sup_lt_sup_of_lt_of_inf_le_inf hAB (le_of_eq h.symm)⟩) - (InvImage.wf _ (h₁.wf.prod_lex h₂.wf))⟩ + StrictMono.wellFoundedLT (f := fun A ↦ (f₂ A, g₂ A)) fun A B hAB ↦ by + simp only [Prod.le_def, lt_iff_le_not_le, ← gci.l_le_l_iff, ← gi.u_le_u_iff, hf, hg] + exact strictMono_inf_prod_sup hAB /-- A generalization of the theorem that if `N` is a submodule of `M` and `N` and `M / N` are both Noetherian, then `M` is Noetherian. -/ @@ -261,7 +258,7 @@ def infIccOrderIsoIccSup (a b : α) : Set.Icc (a ⊓ b) a ≃o Set.Icc b (a ⊔ (by change a ⊓ ↑x ⊔ b = ↑x rw [inf_comm, inf_sup_assoc_of_le _ x.prop.1, inf_eq_left.2 x.prop.2]) - map_rel_iff' := @fun x y => by + map_rel_iff' {x y} := by simp only [Subtype.mk_le_mk, Equiv.coe_fn_mk, le_sup_right] rw [← Subtype.coe_le_coe] refine ⟨fun h => ?_, fun h => sup_le_sup_right h _⟩ From b53b6c31ea3b1e63f6b82642be4e4a72c901668d Mon Sep 17 00:00:00 2001 From: damiano Date: Mon, 7 Oct 2024 11:43:33 +0000 Subject: [PATCH 15/43] chore: remove some unused variables (#17479) --- Mathlib/Algebra/Algebra/Basic.lean | 7 ++-- Mathlib/Algebra/Group/Action/Opposite.lean | 2 +- Mathlib/Algebra/Order/Floor/Div.lean | 2 +- .../Algebra/Order/Monoid/Canonical/Defs.lean | 2 +- Mathlib/Algebra/Order/Sub/Defs.lean | 8 ++--- Mathlib/Algebra/Ring/Action/Basic.lean | 6 ++-- Mathlib/Algebra/Ring/Units.lean | 4 +-- Mathlib/Analysis/Normed/Group/Seminorm.lean | 20 +++++------ .../ConcreteCategory/Bundled.lean | 2 +- .../Enumerative/DoubleCounting.lean | 8 ++--- Mathlib/Computability/Primrec.lean | 13 ++++--- Mathlib/Control/Lawful.lean | 1 - Mathlib/Control/ULift.lean | 2 +- Mathlib/Data/List/Monad.lean | 4 +-- Mathlib/LinearAlgebra/BilinearMap.lean | 18 ++++------ .../RingTheory/NonUnitalSubring/Basic.lean | 36 ++++++------------- Mathlib/Topology/Constructions.lean | 10 +++--- 17 files changed, 60 insertions(+), 85 deletions(-) diff --git a/Mathlib/Algebra/Algebra/Basic.lean b/Mathlib/Algebra/Algebra/Basic.lean index a7a16de103ee9..f4a24f5977006 100644 --- a/Mathlib/Algebra/Algebra/Basic.lean +++ b/Mathlib/Algebra/Algebra/Basic.lean @@ -23,12 +23,12 @@ universe u v w u₁ v₁ namespace Algebra -variable {R : Type u} {S : Type v} {A : Type w} {B : Type*} +variable {R : Type u} {A : Type w} section Semiring -variable [CommSemiring R] [CommSemiring S] -variable [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] +variable [CommSemiring R] +variable [Semiring A] [Algebra R A] section PUnit @@ -319,7 +319,6 @@ section IsScalarTower variable {R : Type*} [CommSemiring R] variable (A : Type*) [Semiring A] [Algebra R A] variable {M : Type*} [AddCommMonoid M] [Module A M] [Module R M] [IsScalarTower R A M] -variable {N : Type*} [AddCommMonoid N] [Module A N] [Module R N] [IsScalarTower R A N] theorem algebra_compatible_smul (r : R) (m : M) : r • m = (algebraMap R A) r • m := by rw [← one_smul A m, ← smul_assoc, Algebra.smul_def, mul_one, one_smul] diff --git a/Mathlib/Algebra/Group/Action/Opposite.lean b/Mathlib/Algebra/Group/Action/Opposite.lean index 8092a68ff91c7..252781fd7b4bd 100644 --- a/Mathlib/Algebra/Group/Action/Opposite.lean +++ b/Mathlib/Algebra/Group/Action/Opposite.lean @@ -27,7 +27,7 @@ With `open scoped RightActions`, this provides: assert_not_exists MonoidWithZero -variable {R M N α : Type*} +variable {M N α : Type*} /-! ### Actions _on_ the opposite type diff --git a/Mathlib/Algebra/Order/Floor/Div.lean b/Mathlib/Algebra/Order/Floor/Div.lean index 35d76dbf57197..20b853e6e5f94 100644 --- a/Mathlib/Algebra/Order/Floor/Div.lean +++ b/Mathlib/Algebra/Order/Floor/Div.lean @@ -119,7 +119,7 @@ end OrderedAddCommMonoid section LinearOrderedAddCommMonoid variable [LinearOrderedAddCommMonoid α] [OrderedAddCommMonoid β] [SMulZeroClass α β] - [PosSMulReflectLE α β] [FloorDiv α β] [CeilDiv α β] {a : α} {b c : β} + [PosSMulReflectLE α β] [FloorDiv α β] [CeilDiv α β] {a : α} {b : β} lemma floorDiv_le_ceilDiv : b ⌊/⌋ a ≤ b ⌈/⌉ a := by obtain ha | ha := le_or_lt a 0 diff --git a/Mathlib/Algebra/Order/Monoid/Canonical/Defs.lean b/Mathlib/Algebra/Order/Monoid/Canonical/Defs.lean index 6e2104c91ef78..7cf5e1c4143ca 100644 --- a/Mathlib/Algebra/Order/Monoid/Canonical/Defs.lean +++ b/Mathlib/Algebra/Order/Monoid/Canonical/Defs.lean @@ -58,7 +58,7 @@ instance (priority := 100) CanonicallyOrderedCommMonoid.existsMulOfLE (α : Type section CanonicallyOrderedCommMonoid -variable [CanonicallyOrderedCommMonoid α] {a b c d : α} +variable [CanonicallyOrderedCommMonoid α] {a b c : α} @[to_additive] theorem le_self_mul : a ≤ a * c := diff --git a/Mathlib/Algebra/Order/Sub/Defs.lean b/Mathlib/Algebra/Order/Sub/Defs.lean index 55d910166545c..9966239cd270f 100644 --- a/Mathlib/Algebra/Order/Sub/Defs.lean +++ b/Mathlib/Algebra/Order/Sub/Defs.lean @@ -41,7 +41,7 @@ TODO: generalize `Nat.le_of_le_of_sub_le_sub_right`, `Nat.sub_le_sub_right_iff`, -/ -variable {α β : Type*} +variable {α : Type*} /-- `OrderedSub α` means that `α` has a subtraction characterized by `a - b ≤ c ↔ a ≤ c + b`. In other words, `a - b` is the least `c` such that `a ≤ b + c`. @@ -60,7 +60,7 @@ theorem tsub_le_iff_right [LE α] [Add α] [Sub α] [OrderedSub α] {a b c : α} a - b ≤ c ↔ a ≤ c + b := OrderedSub.tsub_le_iff_right a b c -variable [Preorder α] [Add α] [Sub α] [OrderedSub α] {a b c d : α} +variable [Preorder α] [Add α] [Sub α] [OrderedSub α] {a b : α} /-- See `add_tsub_cancel_right` for the equality if `ContravariantClass α α (+) (≤)`. -/ theorem add_tsub_le_right : a + b - b ≤ a := @@ -210,7 +210,7 @@ end Contra end AddCommSemigroup -variable [AddCommMonoid α] [Sub α] [OrderedSub α] {a b c d : α} +variable [AddCommMonoid α] [Sub α] [OrderedSub α] {a b : α} theorem tsub_nonpos : a - b ≤ 0 ↔ a ≤ b := by rw [tsub_le_iff_left, add_zero] @@ -375,7 +375,7 @@ end OrderedAddCommSemigroup section LinearOrder -variable {a b c d : α} [LinearOrder α] [AddCommSemigroup α] [Sub α] [OrderedSub α] +variable {a b c : α} [LinearOrder α] [AddCommSemigroup α] [Sub α] [OrderedSub α] /-- See `lt_of_tsub_lt_tsub_right_of_le` for a weaker statement in a partial order. -/ theorem lt_of_tsub_lt_tsub_right (h : a - c < b - c) : a < b := diff --git a/Mathlib/Algebra/Ring/Action/Basic.lean b/Mathlib/Algebra/Ring/Action/Basic.lean index 25558ac78f579..d085e72efe2c3 100644 --- a/Mathlib/Algebra/Ring/Action/Basic.lean +++ b/Mathlib/Algebra/Ring/Action/Basic.lean @@ -42,8 +42,8 @@ class MulSemiringAction (M : Type u) (R : Type v) [Monoid M] [Semiring R] extend section Semiring -variable (M N G : Type*) [Monoid M] [Monoid N] [Group G] -variable (A R S F : Type v) [AddMonoid A] [Semiring R] [CommSemiring S] +variable (M N : Type*) [Monoid M] [Monoid N] +variable (R : Type v) [Semiring R] -- note we could not use `extends` since these typeclasses are made with `old_structure_cmd` instance (priority := 100) MulSemiringAction.toMulDistribMulAction [h : MulSemiringAction M R] : @@ -92,8 +92,6 @@ end section SimpLemmas -variable {M G A R F} - attribute [simp] smul_one smul_mul' smul_zero smul_add end SimpLemmas diff --git a/Mathlib/Algebra/Ring/Units.lean b/Mathlib/Algebra/Ring/Units.lean index 3e177e160d871..619def14cb01a 100644 --- a/Mathlib/Algebra/Ring/Units.lean +++ b/Mathlib/Algebra/Ring/Units.lean @@ -15,7 +15,7 @@ import Mathlib.Algebra.Ring.Hom.Defs universe u v w x -variable {α : Type u} {β : Type v} {γ : Type w} {R : Type x} +variable {α : Type u} {β : Type v} {R : Type x} open Function @@ -23,7 +23,7 @@ namespace Units section HasDistribNeg -variable [Monoid α] [HasDistribNeg α] {a b : α} +variable [Monoid α] [HasDistribNeg α] /-- Each element of the group of units of a ring has an additive inverse. -/ instance : Neg αˣ := diff --git a/Mathlib/Analysis/Normed/Group/Seminorm.lean b/Mathlib/Analysis/Normed/Group/Seminorm.lean index ca8445c51637b..06d80361d16b3 100644 --- a/Mathlib/Analysis/Normed/Group/Seminorm.lean +++ b/Mathlib/Analysis/Normed/Group/Seminorm.lean @@ -48,7 +48,7 @@ open Set open NNReal -variable {ι R R' E F G : Type*} +variable {R R' E F G : Type*} /-- A seminorm on an additive group `G` is a function `f : G → ℝ` that preserves zero, is subadditive and such that `f (-x) = f x` for all `x`. -/ @@ -325,7 +325,7 @@ end Group section CommGroup -variable [CommGroup E] [CommGroup F] (p q : GroupSeminorm E) (x y : E) +variable [CommGroup E] [CommGroup F] (p q : GroupSeminorm E) (x : E) @[to_additive] theorem comp_mul_le (f g : F →* E) : p.comp (f * g) ≤ p.comp f + p.comp g := fun _ => @@ -381,7 +381,7 @@ end GroupSeminorm see that `SMul R ℝ` should be fixed because `ℝ` is fixed. -/ namespace AddGroupSeminorm -variable [AddGroup E] [SMul R ℝ] [SMul R ℝ≥0] [IsScalarTower R ℝ≥0 ℝ] (p : AddGroupSeminorm E) +variable [AddGroup E] [SMul R ℝ] [SMul R ℝ≥0] [IsScalarTower R ℝ≥0 ℝ] instance toOne [DecidableEq E] : One (AddGroupSeminorm E) := ⟨{ toFun := fun x => if x = 0 then 0 else 1 @@ -435,7 +435,7 @@ namespace NonarchAddGroupSeminorm section AddGroup -variable [AddGroup E] [AddGroup F] [AddGroup G] {p q : NonarchAddGroupSeminorm E} +variable [AddGroup E] {p q : NonarchAddGroupSeminorm E} instance funLike : FunLike (NonarchAddGroupSeminorm E) E ℝ where coe f := f.toFun @@ -477,7 +477,7 @@ theorem coe_le_coe : (p : E → ℝ) ≤ q ↔ p ≤ q := theorem coe_lt_coe : (p : E → ℝ) < q ↔ p < q := Iff.rfl -variable (p q) (f : F →+ E) +variable (p q) instance : Zero (NonarchAddGroupSeminorm E) := ⟨{ toFun := 0 @@ -522,7 +522,7 @@ end AddGroup section AddCommGroup -variable [AddCommGroup E] [AddCommGroup F] (p q : NonarchAddGroupSeminorm E) (x y : E) +variable [AddCommGroup E] theorem add_bddBelow_range_add {p q : NonarchAddGroupSeminorm E} {x : E} : BddBelow (range fun y => p y + q (x - y)) := @@ -653,7 +653,7 @@ namespace GroupNorm section Group -variable [Group E] [Group F] [Group G] {p q : GroupNorm E} +variable [Group E] {p q : GroupNorm E} @[to_additive] instance funLike : FunLike (GroupNorm E) E ℝ where @@ -703,7 +703,7 @@ theorem coe_le_coe : (p : E → ℝ) ≤ q ↔ p ≤ q := theorem coe_lt_coe : (p : E → ℝ) < q ↔ p < q := Iff.rfl -variable (p q) (f : F →* E) +variable (p q) @[to_additive] instance : Add (GroupNorm E) := @@ -787,7 +787,7 @@ namespace NonarchAddGroupNorm section AddGroup -variable [AddGroup E] [AddGroup F] {p q : NonarchAddGroupNorm E} +variable [AddGroup E] {p q : NonarchAddGroupNorm E} instance funLike : FunLike (NonarchAddGroupNorm E) E ℝ where coe f := f.toFun @@ -829,7 +829,7 @@ theorem coe_le_coe : (p : E → ℝ) ≤ q ↔ p ≤ q := theorem coe_lt_coe : (p : E → ℝ) < q ↔ p < q := Iff.rfl -variable (p q) (f : F →+ E) +variable (p q) instance : Sup (NonarchAddGroupNorm E) := ⟨fun p q => diff --git a/Mathlib/CategoryTheory/ConcreteCategory/Bundled.lean b/Mathlib/CategoryTheory/ConcreteCategory/Bundled.lean index d2a6dab1aa44a..1fe01ebf2d14a 100644 --- a/Mathlib/CategoryTheory/ConcreteCategory/Bundled.lean +++ b/Mathlib/CategoryTheory/ConcreteCategory/Bundled.lean @@ -22,7 +22,7 @@ universe u v namespace CategoryTheory -variable {c d : Type u → Type v} {α : Type u} +variable {c d : Type u → Type v} /-- `Bundled` is a type bundled with a type class instance for that type. Only the type class is exposed as a parameter. -/ diff --git a/Mathlib/Combinatorics/Enumerative/DoubleCounting.lean b/Mathlib/Combinatorics/Enumerative/DoubleCounting.lean index 542deff1911e1..0264391a00385 100644 --- a/Mathlib/Combinatorics/Enumerative/DoubleCounting.lean +++ b/Mathlib/Combinatorics/Enumerative/DoubleCounting.lean @@ -39,7 +39,7 @@ namespace Finset section Bipartite -variable (r : α → β → Prop) (s : Finset α) (t : Finset β) (a a' : α) (b b' : β) +variable (r : α → β → Prop) (s : Finset α) (t : Finset β) (a : α) (b : β) [DecidablePred (r a)] [∀ a, Decidable (r a b)] {m n : ℕ} /-- Elements of `s` which are "below" `b` according to relation `r`. -/ @@ -58,7 +58,7 @@ theorem coe_bipartiteBelow : s.bipartiteBelow r b = ({a ∈ s | r a b} : Set α) @[simp, norm_cast] theorem coe_bipartiteAbove : t.bipartiteAbove r a = ({b ∈ t | r a b} : Set β) := coe_filter _ _ -variable {s t a a' b b'} +variable {s t a b} @[simp] theorem mem_bipartiteBelow {a : α} : a ∈ s.bipartiteBelow r b ↔ a ∈ s ∧ r a b := mem_filter @@ -72,7 +72,7 @@ theorem sum_card_bipartiteAbove_eq_sum_card_bipartiteBelow [∀ a b, Decidable ( exact sum_comm section OrderedSemiring -variable [OrderedSemiring R] [DecidablePred (r a)] [∀ a, Decidable (r a b)] {m n : R} +variable [OrderedSemiring R] {m n : R} /-- **Double counting** argument. @@ -100,7 +100,7 @@ end OrderedSemiring section StrictOrderedSemiring variable [StrictOrderedSemiring R] (r : α → β → Prop) {s : Finset α} {t : Finset β} - (a a' : α) (b b' : β) [DecidablePred (r a)] [∀ a, Decidable (r a b)] {m n : R} + (a b) {m n : R} /-- **Double counting** argument. diff --git a/Mathlib/Computability/Primrec.lean b/Mathlib/Computability/Primrec.lean index 103424afc40a5..350ee120e6376 100644 --- a/Mathlib/Computability/Primrec.lean +++ b/Mathlib/Computability/Primrec.lean @@ -275,7 +275,7 @@ end Primcodable namespace Primrec -variable {α : Type*} {σ : Type*} [Primcodable α] [Primcodable σ] +variable {α : Type*} [Primcodable α] open Nat.Primrec @@ -457,8 +457,8 @@ end Primrec₂ namespace Primrec -variable {α : Type*} {β : Type*} {γ : Type*} {δ : Type*} {σ : Type*} -variable [Primcodable α] [Primcodable β] [Primcodable γ] [Primcodable δ] [Primcodable σ] +variable {α : Type*} {β : Type*} {σ : Type*} +variable [Primcodable α] [Primcodable β] [Primcodable σ] theorem to₂ {f : α × β → σ} (hf : Primrec f) : Primrec₂ fun a b => f (a, b) := hf.of_eq fun _ => rfl @@ -1088,8 +1088,7 @@ end Primrec namespace Primcodable -variable {α : Type*} {β : Type*} -variable [Primcodable α] [Primcodable β] +variable {α : Type*} [Primcodable α] open Primrec @@ -1139,8 +1138,8 @@ end Primcodable namespace Primrec -variable {α : Type*} {β : Type*} {γ : Type*} {σ : Type*} -variable [Primcodable α] [Primcodable β] [Primcodable γ] [Primcodable σ] +variable {α : Type*} {β : Type*} {σ : Type*} +variable [Primcodable α] [Primcodable β] [Primcodable σ] theorem subtype_val {p : α → Prop} [DecidablePred p] {hp : PrimrecPred p} : haveI := Primcodable.subtype hp diff --git a/Mathlib/Control/Lawful.lean b/Mathlib/Control/Lawful.lean index ce4807b4cde66..f2cd84ab5da94 100644 --- a/Mathlib/Control/Lawful.lean +++ b/Mathlib/Control/Lawful.lean @@ -73,7 +73,6 @@ namespace ReaderT section -variable {ρ : Type u} variable {m : Type u → Type v} variable {α σ : Type u} diff --git a/Mathlib/Control/ULift.lean b/Mathlib/Control/ULift.lean index f294e51cf20de..2482e0095fe06 100644 --- a/Mathlib/Control/ULift.lean +++ b/Mathlib/Control/ULift.lean @@ -78,7 +78,7 @@ end PLift namespace ULift -variable {α : Type u} {β : Type v} {f : α → β} +variable {α : Type u} {β : Type v} /-- Functorial action. -/ protected def map (f : α → β) (a : ULift.{u'} α) : ULift.{v'} β := ULift.up.{v'} (f a.down) diff --git a/Mathlib/Data/List/Monad.lean b/Mathlib/Data/List/Monad.lean index a89b5f05d5272..04d1477cbd14d 100644 --- a/Mathlib/Data/List/Monad.lean +++ b/Mathlib/Data/List/Monad.lean @@ -9,11 +9,11 @@ import Mathlib.Init # Monad instances for `List` -/ -universe u v w +universe u namespace List -variable {α : Type u} {β : Type v} {γ : Type w} +variable {α : Type u} instance instMonad : Monad List.{u} where pure := @List.pure diff --git a/Mathlib/LinearAlgebra/BilinearMap.lean b/Mathlib/LinearAlgebra/BilinearMap.lean index b9e49faf0456c..10c008423626e 100644 --- a/Mathlib/LinearAlgebra/BilinearMap.lean +++ b/Mathlib/LinearAlgebra/BilinearMap.lean @@ -37,16 +37,15 @@ variable {R : Type*} [Semiring R] {S : Type*} [Semiring S] variable {R₂ : Type*} [Semiring R₂] {S₂ : Type*} [Semiring S₂] variable {M : Type*} {N : Type*} {P : Type*} variable {M₂ : Type*} {N₂ : Type*} {P₂ : Type*} -variable {Nₗ : Type*} {Pₗ : Type*} -variable {M' : Type*} {N' : Type*} {P' : Type*} +variable {Pₗ : Type*} +variable {M' : Type*} {P' : Type*} variable [AddCommMonoid M] [AddCommMonoid N] [AddCommMonoid P] -variable [AddCommMonoid M₂] [AddCommMonoid N₂] [AddCommMonoid P₂] -variable [AddCommMonoid Nₗ] [AddCommMonoid Pₗ] -variable [AddCommGroup M'] [AddCommGroup N'] [AddCommGroup P'] +variable [AddCommMonoid M₂] [AddCommMonoid N₂] [AddCommMonoid P₂] [AddCommMonoid Pₗ] +variable [AddCommGroup M'] [AddCommGroup P'] variable [Module R M] [Module S N] [Module R₂ P] [Module S₂ P] variable [Module R M₂] [Module S N₂] [Module R P₂] [Module S₂ P₂] variable [Module R Pₗ] [Module S Pₗ] -variable [Module R M'] [Module S N'] [Module R₂ P'] [Module S₂ P'] +variable [Module R M'] [Module R₂ P'] [Module S₂ P'] variable [SMulCommClass S₂ R₂ P] [SMulCommClass S R Pₗ] [SMulCommClass S₂ R₂ P'] variable [SMulCommClass S₂ R P₂] variable {ρ₁₂ : R →+* R₂} {σ₁₂ : S →+* S₂} @@ -376,14 +375,11 @@ end CommSemiring section CommRing -variable {R R₂ S S₂ M N P : Type*} -variable {Mₗ Nₗ Pₗ : Type*} -variable [CommRing R] [CommRing S] [CommRing R₂] [CommRing S₂] +variable {R M : Type*} [CommRing R] section AddCommGroup -variable [AddCommGroup M] [AddCommGroup N] [AddCommGroup P] -variable [Module R M] [Module S N] [Module R₂ P] [Module S₂ P] +variable [AddCommGroup M] [Module R M] theorem lsmul_injective [NoZeroSMulDivisors R M] {x : R} (hx : x ≠ 0) : Function.Injective (lsmul R M x) := diff --git a/Mathlib/RingTheory/NonUnitalSubring/Basic.lean b/Mathlib/RingTheory/NonUnitalSubring/Basic.lean index 69c9fcd2931b4..f71dbda3f4e85 100644 --- a/Mathlib/RingTheory/NonUnitalSubring/Basic.lean +++ b/Mathlib/RingTheory/NonUnitalSubring/Basic.lean @@ -71,7 +71,7 @@ universe u v w section Basic -variable {R : Type u} {S : Type v} {T : Type w} [NonUnitalNonAssocRing R] +variable {R : Type u} {S : Type v} [NonUnitalNonAssocRing R] section NonUnitalSubringClass @@ -124,8 +124,6 @@ end NonUnitalSubringClass end NonUnitalSubringClass -variable [NonUnitalNonAssocRing S] [NonUnitalNonAssocRing T] - /-- `NonUnitalSubring R` is the type of non-unital subrings of `R`. A non-unital subring of `R` is a subset `s` that is a multiplicative subsemigroup and an additive subgroup. Note in particular that it shares the same 0 as R. -/ @@ -387,7 +385,7 @@ section Hom namespace NonUnitalSubring -variable {F : Type w} {R : Type u} {S : Type v} {T : Type*} {SR : Type*} +variable {F : Type w} {R : Type u} {S : Type v} {T : Type*} [NonUnitalNonAssocRing R] [NonUnitalNonAssocRing S] [NonUnitalNonAssocRing T] [FunLike F R S] [NonUnitalRingHomClass F R S] (s : NonUnitalSubring R) @@ -503,10 +501,7 @@ namespace NonUnitalSubring section Order -variable {F : Type w} {R : Type u} {S : Type v} {T : Type*} - [NonUnitalNonAssocRing R] [NonUnitalNonAssocRing S] [NonUnitalNonAssocRing T] - [FunLike F R S] [NonUnitalRingHomClass F R S] - (g : S →ₙ+* T) (f : R →ₙ+* S) +variable {R : Type u} [NonUnitalNonAssocRing R] /-! ## bot -/ @@ -640,10 +635,9 @@ end Center /-! ## `NonUnitalSubring` closure of a subset -/ -variable {F : Type w} {R : Type u} {S : Type v} {T : Type*} - [NonUnitalNonAssocRing R] [NonUnitalNonAssocRing S] [NonUnitalNonAssocRing T] +variable {F : Type w} {R : Type u} {S : Type v} + [NonUnitalNonAssocRing R] [NonUnitalNonAssocRing S] [FunLike F R S] [NonUnitalRingHomClass F R S] - (g : S →ₙ+* T) (f : R →ₙ+* S) /-- The `NonUnitalSubring` generated by a set. -/ def closure (s : Set R) : NonUnitalSubring R := @@ -910,11 +904,8 @@ end NonUnitalSubring namespace NonUnitalRingHom -variable {F : Type w} {R : Type u} {S : Type v} {T : Type*} - [NonUnitalNonAssocRing R] [NonUnitalNonAssocRing S] [NonUnitalNonAssocRing T] - [FunLike F R S] [NonUnitalRingHomClass F R S] - (g : S →ₙ+* T) (f : R →ₙ+* S) - {s : NonUnitalSubring R} +variable {R : Type u} {S : Type v} + [NonUnitalNonAssocRing R] [NonUnitalNonAssocRing S] open NonUnitalSubring @@ -980,11 +971,8 @@ end NonUnitalRingHom namespace NonUnitalSubring -variable {F : Type w} {R : Type u} {S : Type v} {T : Type*} - [NonUnitalNonAssocRing R] [NonUnitalNonAssocRing S] [NonUnitalNonAssocRing T] - [FunLike F R S] [NonUnitalRingHomClass F R S] - (g : S →ₙ+* T) (f : R →ₙ+* S) - {s : NonUnitalSubring R} +variable {R : Type u} {S : Type v} + [NonUnitalNonAssocRing R] [NonUnitalNonAssocRing S] open NonUnitalRingHom @@ -1006,11 +994,7 @@ end NonUnitalSubring namespace RingEquiv -variable {F : Type w} {R : Type u} {S : Type v} {T : Type*} - [NonUnitalRing R] [NonUnitalRing S] [NonUnitalRing T] - [FunLike F R S] [NonUnitalRingHomClass F R S] - (g : S →ₙ+* T) (f : R →ₙ+* S) - {s t : NonUnitalSubring R} +variable {R : Type u} {S : Type v} [NonUnitalRing R] [NonUnitalRing S] {s t : NonUnitalSubring R} /-- Makes the identity isomorphism from a proof two `NonUnitalSubring`s of a multiplicative monoid are equal. -/ diff --git a/Mathlib/Topology/Constructions.lean b/Mathlib/Topology/Constructions.lean index f23b56ce022b3..803aee5a4888b 100644 --- a/Mathlib/Topology/Constructions.lean +++ b/Mathlib/Topology/Constructions.lean @@ -967,7 +967,7 @@ end Sum section Subtype -variable [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] {p : X → Prop} +variable [TopologicalSpace X] [TopologicalSpace Y] {p : X → Prop} theorem inducing_subtype_val {t : Set Y} : Inducing ((↑) : t → Y) := ⟨rfl⟩ @@ -1122,7 +1122,7 @@ end Subtype section Quotient -variable [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] +variable [TopologicalSpace X] [TopologicalSpace Y] variable {r : X → X → Prop} {s : Setoid X} theorem quotientMap_quot_mk : QuotientMap (@Quot.mk X r) := @@ -1248,7 +1248,7 @@ theorem Finset.continuous_restrict₂ {s t : Finset ι} (hst : s ⊆ t) : Continuous (Finset.restrict₂ (π := π) hst) := continuous_pi fun _ ↦ continuous_apply _ -variable {Z : Type*} [TopologicalSpace Z] +variable [TopologicalSpace Z] @[continuity, fun_prop] theorem Pi.continuous_restrict_apply (s : Set X) {f : X → Z} (hf : Continuous f) : @@ -1642,8 +1642,8 @@ theorem IsClosed.trans (ht : IsClosed t) (hs : IsClosed s) : IsClosed (t : Set X end Monad section NhdsSet -variable {X Y : Type*} [TopologicalSpace X] [TopologicalSpace Y] {f : Filter X} - {s : Set X} {t : Set Y} {x : X} +variable [TopologicalSpace X] [TopologicalSpace Y] + {s : Set X} {t : Set Y} /-- The product of a neighborhood of `s` and a neighborhood of `t` is a neighborhood of `s ×ˢ t`, formulated in terms of a filter inequality. -/ From 78b3a02235527543f0c61118f06b48d44fa96a32 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 7 Oct 2024 12:39:02 +0000 Subject: [PATCH 16/43] fix: resolve ofNat vs natCast confusion in a Fin lemma (#17422) --- Mathlib/Data/Fin/Basic.lean | 8 +++++++- Mathlib/Logic/Equiv/Fin.lean | 2 +- 2 files changed, 8 insertions(+), 2 deletions(-) diff --git a/Mathlib/Data/Fin/Basic.lean b/Mathlib/Data/Fin/Basic.lean index d875809cf1a4a..e0de8283c150d 100644 --- a/Mathlib/Data/Fin/Basic.lean +++ b/Mathlib/Data/Fin/Basic.lean @@ -1514,10 +1514,16 @@ lemma sub_succ_le_sub_of_le {n : ℕ} {u v : Fin (n + 2)} (h : u < v) : v - (u + end AddGroup @[simp] -theorem coe_ofNat_eq_mod (m n : ℕ) [NeZero m] : +theorem coe_natCast_eq_mod (m n : ℕ) [NeZero m] : ((n : Fin m) : ℕ) = n % m := rfl +-- See note [no_index around OfNat.ofNat] +@[simp] +theorem coe_ofNat_eq_mod (m n : ℕ) [NeZero m] : + ((no_index OfNat.ofNat n : Fin m) : ℕ) = OfNat.ofNat n % m := + rfl + section Mul /-! diff --git a/Mathlib/Logic/Equiv/Fin.lean b/Mathlib/Logic/Equiv/Fin.lean index 1d07d2d1b2b25..3453d19eac020 100644 --- a/Mathlib/Logic/Equiv/Fin.lean +++ b/Mathlib/Logic/Equiv/Fin.lean @@ -470,7 +470,7 @@ def Int.divModEquiv (n : ℕ) [NeZero n] : ℤ ≃ ℤ × Fin n where toFun a := (a / n, ↑(a.natMod n)) invFun p := p.1 * n + ↑p.2 left_inv a := by - simp_rw [Fin.coe_ofNat_eq_mod, natCast_mod, natMod, + simp_rw [Fin.coe_natCast_eq_mod, natCast_mod, natMod, toNat_of_nonneg (emod_nonneg _ <| natCast_eq_zero.not.2 (NeZero.ne n)), emod_emod, ediv_add_emod'] right_inv := fun ⟨q, r, hrn⟩ => by From 685ce1b115cb15b72d0610991eff2433888349a6 Mon Sep 17 00:00:00 2001 From: Bhavik Mehta Date: Mon, 7 Oct 2024 13:31:20 +0000 Subject: [PATCH 17/43] chore(Data/Finset): subset_image_iff about finsets (#17036) State `subset_image_iff` for finsets, and rename the existing set/finset version to a more appropriate name. Also make the analogous change for image2 --- .../Algebra/Group/Pointwise/Finset/Basic.lean | 8 ++++---- Mathlib/Data/Finset/Image.lean | 16 +++++++++++++++- Mathlib/Data/Finset/NAry.lean | 6 ++++-- Mathlib/Data/Finset/Preimage.lean | 2 +- Mathlib/Data/Finset/Sups.lean | 6 +++--- Mathlib/Data/Set/Finite.lean | 2 +- Mathlib/RingTheory/Finiteness.lean | 4 +--- 7 files changed, 29 insertions(+), 15 deletions(-) diff --git a/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean b/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean index 521a3b832afdc..5e5ebeefca3c2 100644 --- a/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean +++ b/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean @@ -420,7 +420,7 @@ theorem union_mul_inter_subset_union : (s₁ ∪ s₂) * (t₁ ∩ t₂) ⊆ s `s'`, `t'` such that `s' ⊆ s`, `t' ⊆ t` and `u ⊆ s' + t'`."] theorem subset_mul {s t : Set α} : ↑u ⊆ s * t → ∃ s' t' : Finset α, ↑s' ⊆ s ∧ ↑t' ⊆ t ∧ u ⊆ s' * t' := - subset_image₂ + subset_set_image₂ @[to_additive] theorem image_mul [DecidableEq β] : (s * t).image (f : α → β) = s.image f * t.image f := @@ -606,7 +606,7 @@ theorem union_div_inter_subset_union : (s₁ ∪ s₂) / (t₁ ∩ t₂) ⊆ s `s'`, `t'` such that `s' ⊆ s`, `t' ⊆ t` and `u ⊆ s' - t'`."] theorem subset_div {s t : Set α} : ↑u ⊆ s / t → ∃ s' t' : Finset α, ↑s' ⊆ s ∧ ↑t' ⊆ t ∧ u ⊆ s' / t' := - subset_image₂ + subset_set_image₂ @[to_additive (attr := simp (default + 1))] lemma sup_div_le [SemilatticeSup β] [OrderBot β] {s t : Finset α} {f : α → β} {a : β} : @@ -1200,7 +1200,7 @@ theorem union_smul_inter_subset_union [DecidableEq α] : (s₁ ∪ s₂) • (t finsets `s'`, `t'` such that `s' ⊆ s`, `t' ⊆ t` and `u ⊆ s' +ᵥ t'`."] theorem subset_smul {s : Set α} {t : Set β} : ↑u ⊆ s • t → ∃ (s' : Finset α) (t' : Finset β), ↑s' ⊆ s ∧ ↑t' ⊆ t ∧ u ⊆ s' • t' := - subset_image₂ + subset_set_image₂ end SMul @@ -1310,7 +1310,7 @@ end finsets `s'`, `t'` such that `s' ⊆ s`, `t' ⊆ t` and `u ⊆ s' -ᵥ t'`. -/ theorem subset_vsub {s t : Set β} : ↑u ⊆ s -ᵥ t → ∃ s' t' : Finset β, ↑s' ⊆ s ∧ ↑t' ⊆ t ∧ u ⊆ s' -ᵥ t' := - subset_image₂ + subset_set_image₂ end VSub diff --git a/Mathlib/Data/Finset/Image.lean b/Mathlib/Data/Finset/Image.lean index 3fc36f5e38e47..88939058da20c 100644 --- a/Mathlib/Data/Finset/Image.lean +++ b/Mathlib/Data/Finset/Image.lean @@ -697,9 +697,23 @@ theorem fin_mono {n} : Monotone (Finset.fin n) := fun s t h x => by simpa using theorem fin_map {n} {s : Finset ℕ} : (s.fin n).map Fin.valEmbedding = s.filter (· < n) := by simp [Finset.fin, Finset.map_map] +/-- +If a finset `t` is a subset of the image of another finset `s` under `f`, then it is equal to the +image of a subset of `s`. + +For the version where `s` is a set, see `subset_set_image_iff`. +-/ +theorem subset_image_iff [DecidableEq β] {s : Finset α} {t : Finset β} {f : α → β} : + t ⊆ s.image f ↔ ∃ s' : Finset α, s' ⊆ s ∧ s'.image f = t := by + refine ⟨fun ht => ?_, fun ⟨s', hs', h⟩ => h ▸ image_subset_image hs'⟩ + refine ⟨s.filter (f · ∈ t), filter_subset _ _, le_antisymm (by simp [image_subset_iff]) ?_⟩ + intro x hx + specialize ht hx + aesop + /-- If a `Finset` is a subset of the image of a `Set` under `f`, then it is equal to the `Finset.image` of a `Finset` subset of that `Set`. -/ -theorem subset_image_iff [DecidableEq β] {s : Set α} {t : Finset β} {f : α → β} : +theorem subset_set_image_iff [DecidableEq β] {s : Set α} {t : Finset β} {f : α → β} : ↑t ⊆ f '' s ↔ ∃ s' : Finset α, ↑s' ⊆ s ∧ s'.image f = t := by constructor; swap · rintro ⟨t, ht, rfl⟩ diff --git a/Mathlib/Data/Finset/NAry.lean b/Mathlib/Data/Finset/NAry.lean index fa32e7a832096..4db70840b348d 100644 --- a/Mathlib/Data/Finset/NAry.lean +++ b/Mathlib/Data/Finset/NAry.lean @@ -454,9 +454,9 @@ theorem card_dvd_card_image₂_left (hf : ∀ b ∈ t, Injective fun a => f a b) /-- If a `Finset` is a subset of the image of two `Set`s under a binary operation, then it is a subset of the `Finset.image₂` of two `Finset` subsets of these `Set`s. -/ -theorem subset_image₂ {s : Set α} {t : Set β} (hu : ↑u ⊆ image2 f s t) : +theorem subset_set_image₂ {s : Set α} {t : Set β} (hu : ↑u ⊆ image2 f s t) : ∃ (s' : Finset α) (t' : Finset β), ↑s' ⊆ s ∧ ↑t' ⊆ t ∧ u ⊆ image₂ f s' t' := by - rw [← Set.image_prod, subset_image_iff] at hu + rw [← Set.image_prod, subset_set_image_iff] at hu rcases hu with ⟨u, hu, rfl⟩ classical use u.image Prod.fst, u.image Prod.snd @@ -464,6 +464,8 @@ theorem subset_image₂ {s : Set α} {t : Set β} (hu : ↑u ⊆ image2 f s t) : image_subset_iff] exact ⟨fun _ h ↦ (hu h).1, fun _ h ↦ (hu h).2, fun x hx ↦ mem_image₂_of_mem hx hx⟩ +@[deprecated (since := "2024-09-22")] alias subset_image₂ := subset_set_image₂ + end section UnionInter diff --git a/Mathlib/Data/Finset/Preimage.lean b/Mathlib/Data/Finset/Preimage.lean index 45d7cdc8ae557..657acee065919 100644 --- a/Mathlib/Data/Finset/Preimage.lean +++ b/Mathlib/Data/Finset/Preimage.lean @@ -100,7 +100,7 @@ theorem preimage_subset {f : α ↪ β} {s : Finset β} {t : Finset α} (hs : s theorem subset_map_iff {f : α ↪ β} {s : Finset β} {t : Finset α} : s ⊆ t.map f ↔ ∃ u ⊆ t, s = u.map f := by classical - simp_rw [← coe_subset, coe_map, subset_image_iff, map_eq_image, eq_comm] + simp_rw [map_eq_image, subset_image_iff, eq_comm] theorem sigma_preimage_mk {β : α → Type*} [DecidableEq α] (s : Finset (Σa, β a)) (t : Finset α) : (t.sigma fun a => s.preimage (Sigma.mk a) sigma_mk_injective.injOn) = diff --git a/Mathlib/Data/Finset/Sups.lean b/Mathlib/Data/Finset/Sups.lean index 16e426c59c332..5188385616bb7 100644 --- a/Mathlib/Data/Finset/Sups.lean +++ b/Mathlib/Data/Finset/Sups.lean @@ -146,7 +146,7 @@ theorem sups_inter_subset_right : s ⊻ (t₁ ∩ t₂) ⊆ s ⊻ t₁ ∩ s ⊻ theorem subset_sups {s t : Set α} : ↑u ⊆ s ⊻ t → ∃ s' t' : Finset α, ↑s' ⊆ s ∧ ↑t' ⊆ t ∧ u ⊆ s' ⊻ t' := - subset_image₂ + subset_set_image₂ lemma image_sups (f : F) (s t : Finset α) : image f (s ⊻ t) = image f s ⊻ image f t := image_image₂_distrib <| map_sup f @@ -293,7 +293,7 @@ theorem infs_inter_subset_right : s ⊼ (t₁ ∩ t₂) ⊆ s ⊼ t₁ ∩ s ⊼ theorem subset_infs {s t : Set α} : ↑u ⊆ s ⊼ t → ∃ s' t' : Finset α, ↑s' ⊆ s ∧ ↑t' ⊆ t ∧ u ⊆ s' ⊼ t' := - subset_image₂ + subset_set_image₂ lemma image_infs (f : F) (s t : Finset α) : image f (s ⊼ t) = image f s ⊼ image f t := image_image₂_distrib <| map_inf f @@ -582,7 +582,7 @@ lemma diffs_inter_subset_right : s \\ (t₁ ∩ t₂) ⊆ s \\ t₁ ∩ s \\ t lemma subset_diffs {s t : Set α} : ↑u ⊆ Set.image2 (· \ ·) s t → ∃ s' t' : Finset α, ↑s' ⊆ s ∧ ↑t' ⊆ t ∧ u ⊆ s' \\ t' := - subset_image₂ + subset_set_image₂ variable (s t u) diff --git a/Mathlib/Data/Set/Finite.lean b/Mathlib/Data/Set/Finite.lean index dcd9b53e9dcc2..cf7a2bfbb02d2 100644 --- a/Mathlib/Data/Set/Finite.lean +++ b/Mathlib/Data/Set/Finite.lean @@ -897,7 +897,7 @@ theorem exists_subset_image_finite_and {f : α → β} {s : Set α} {p : Set β (∃ t ⊆ f '' s, t.Finite ∧ p t) ↔ ∃ t ⊆ s, t.Finite ∧ p (f '' t) := by classical simp_rw [@and_comm (_ ⊆ _), and_assoc, exists_finite_iff_finset, @and_comm (p _), - Finset.subset_image_iff] + Finset.subset_set_image_iff] aesop section Pi diff --git a/Mathlib/RingTheory/Finiteness.lean b/Mathlib/RingTheory/Finiteness.lean index 19651b5ecbdf0..0039e1437e417 100644 --- a/Mathlib/RingTheory/Finiteness.lean +++ b/Mathlib/RingTheory/Finiteness.lean @@ -381,9 +381,7 @@ theorem fg_iff_compact (s : Submodule R M) : s.FG ↔ CompleteLattice.IsCompactE suffices u.sup id ≤ s from le_antisymm husup this rw [sSup', Finset.sup_id_eq_sSup] exact sSup_le_sSup huspan - -- Porting note: had to split this out of the `obtain` - have := Finset.subset_image_iff.mp huspan - obtain ⟨t, ⟨-, rfl⟩⟩ := this + obtain ⟨t, -, rfl⟩ := Finset.subset_set_image_iff.mp huspan rw [Finset.sup_image, Function.id_comp, Finset.sup_eq_iSup, supr_rw, ← span_eq_iSup_of_singleton_spans, eq_comm] at ssup exact ⟨t, ssup⟩ From 40436c45c46e4bf7d58d660fafd9472d758b9b5d Mon Sep 17 00:00:00 2001 From: Dagur Asgeirsson Date: Mon, 7 Oct 2024 13:31:21 +0000 Subject: [PATCH 18/43] chore(CategoryTheory/Adjunction): deduplicate API for uniqueness of adjoints (#17113) Removes some API from the file `Adjunction.Unique` which was a duplicate of API from `Adjunction.Mates` Co-authored-by: @emilyriehl Co-authored-by: dagurtomas --- Mathlib/CategoryTheory/Adjunction/Basic.lean | 44 ++++++- Mathlib/CategoryTheory/Adjunction/Mates.lean | 2 + Mathlib/CategoryTheory/Adjunction/Unique.lean | 109 ++---------------- 3 files changed, 55 insertions(+), 100 deletions(-) diff --git a/Mathlib/CategoryTheory/Adjunction/Basic.lean b/Mathlib/CategoryTheory/Adjunction/Basic.lean index 30d3ab3bb1913..3d9c1506b6275 100644 --- a/Mathlib/CategoryTheory/Adjunction/Basic.lean +++ b/Mathlib/CategoryTheory/Adjunction/Basic.lean @@ -14,7 +14,9 @@ import Mathlib.CategoryTheory.Yoneda We provide various useful constructors: * `mkOfHomEquiv` -* `mkOfUnitCounit` +* `mk'`: construct an adjunction from the data of a hom set equivalence, unit and counit natural + transformations together with proofs of the equalities `homEquiv_unit` and `homEquiv_counit` + relating them to each other. * `leftAdjointOfEquiv` / `rightAdjointOfEquiv` construct a left/right adjoint of a given functor given the action on objects and the relevant equivalence of morphism spaces. @@ -30,6 +32,44 @@ adjoint can be obtained as `F.rightAdjoint`. `toEquivalence` upgrades an adjunction to an equivalence, given witnesses that the unit and counit are pointwise isomorphisms. Conversely `Equivalence.toAdjunction` recovers the underlying adjunction from an equivalence. + +## Overview of the directory `CategoryTheory.Adjunction` + +* Adjoint lifting theorems are in the directory `Lifting`. +* The file `AdjointFunctorTheorems` proves the adjoint functor theorems. +* The file `Comma` shows that for a functor `G : D ⥤ C` the data of an initial object in each + `StructuredArrow` category on `G` is equivalent to a left adjoint to `G`, as well as the dual. +* The file `Evaluation` shows that products and coproducts are adjoint to evaluation of functors. +* The file `FullyFaithful` characterizes when adjoints are full or faithful in terms of the unit + and counit. +* The file `Limits` proves that left adjoints preserve colimits and right adjoints preserve limits. +* The file `Mates` establishes the bijection between the 2-cells + ``` + L₁ R₁ + C --→ D C ←-- D + G ↓ ↗ ↓ H G ↓ ↘ ↓ H + E --→ F E ←-- F + L₂ R₂ + ``` + where `L₁ ⊣ R₁` and `L₂ ⊣ R₂`. Specializing to a pair of adjoints `L₁ L₂ : C ⥤ D`, + `R₁ R₂ : D ⥤ C`, it provides equivalences `(L₂ ⟶ L₁) ≃ (R₁ ⟶ R₂)` and `(L₂ ≅ L₁) ≃ (R₁ ≅ R₂)`. +* The file `Opposites` contains constructions to relate adjunctions of functors to adjunctions of + their opposites. +* The file `Reflective` defines reflective functors, i.e. fully faithful right adjoints. Note that + many facts about reflective functors are proved in the earlier file `FullyFaithful`. +* The file `Restrict` defines the restriction of an adjunction along fully faithful functors. +* The file `Triple` proves that in an adjoint triple, the left adjoint is fully faithful if and + only if the right adjoint is. +* The file `Unique` proves uniqueness of adjoints. +* The file `Whiskering` proves that functors `F : D ⥤ E` and `G : E ⥤ D` with an adjunction + `F ⊣ G`, induce adjunctions between the functor categories `C ⥤ D` and `C ⥤ E`, + and the functor categories `E ⥤ C` and `D ⥤ C`. + +## Other files related to adjunctions + +* The file `CategoryTheory.Monad.Adjunction` develops the basic relationship between adjunctions + and (co)monads. There it is also shown that given an adjunction `L ⊣ R` and an isomorphism + `L ⋙ R ≅ 𝟭 C`, the unit is an isomorphism, and similarly for the counit. -/ @@ -59,8 +99,6 @@ hom set equivalence. To construct adjoints to a given functor, there are constructors `leftAdjointOfEquiv` and `adjunctionOfEquivLeft` (as well as their duals). -Uniqueness of adjoints is shown in `CategoryTheory.Adjunction.Unique`. - See . -/ structure Adjunction (F : C ⥤ D) (G : D ⥤ C) where diff --git a/Mathlib/CategoryTheory/Adjunction/Mates.lean b/Mathlib/CategoryTheory/Adjunction/Mates.lean index fa2243dda407c..ca344b3aa3796 100644 --- a/Mathlib/CategoryTheory/Adjunction/Mates.lean +++ b/Mathlib/CategoryTheory/Adjunction/Mates.lean @@ -13,11 +13,13 @@ import Mathlib.Tactic.ApplyFun This file establishes the bijection between the 2-cells +``` L₁ R₁ C --→ D C ←-- D G ↓ ↗ ↓ H G ↓ ↘ ↓ H E --→ F E ←-- F L₂ R₂ +``` where `L₁ ⊣ R₁` and `L₂ ⊣ R₂`. The corresponding natural transformations are called mates. diff --git a/Mathlib/CategoryTheory/Adjunction/Unique.lean b/Mathlib/CategoryTheory/Adjunction/Unique.lean index e911b3e4dcb59..9c4b07886fc7d 100644 --- a/Mathlib/CategoryTheory/Adjunction/Unique.lean +++ b/Mathlib/CategoryTheory/Adjunction/Unique.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Bhavik Mehta. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Bhavik Mehta, Thomas Read, Andrew Yang, Dagur Asgeirsson, Joël Riou -/ -import Mathlib.CategoryTheory.Adjunction.Basic +import Mathlib.CategoryTheory.Adjunction.Mates /-! # Uniqueness of adjoints @@ -11,9 +11,6 @@ import Mathlib.CategoryTheory.Adjunction.Basic This file shows that adjoints are unique up to natural isomorphism. ## Main results -* `Adjunction.natTransEquiv` and `Adjunction.natIsoEquiv` If `F ⊣ G` and `F' ⊣ G'` are adjunctions, - then there are equivalences `(G ⟶ G') ≃ (F' ⟶ F)` and `(G ≅ G') ≃ (F' ≅ F)`. -Everything else is deduced from this: * `Adjunction.leftAdjointUniq` : If `F` and `F'` are both left adjoint to `G`, then they are naturally isomorphic. @@ -21,12 +18,6 @@ Everything else is deduced from this: * `Adjunction.rightAdjointUniq` : If `G` and `G'` are both right adjoint to `F`, then they are naturally isomorphic. -## TODO - -There some overlap with the file `Adjunction.Mates`. In particular, `natTransEquiv` is just a -special case of `mateEquiv`. However, before removing `natTransEquiv`, in favour of `mateEquiv`, -the latter needs some more API lemmas such as `natTransEquiv_apply_app`, `natTransEquiv_id`, etc. -in order to make automation work better in the rest of this file. -/ open CategoryTheory @@ -35,89 +26,9 @@ variable {C D : Type*} [Category C] [Category D] namespace CategoryTheory.Adjunction -/-- -If `F ⊣ G` and `F' ⊣ G'` are adjunctions, then giving a natural transformation `G ⟶ G'` is the -same as giving a natural transformation `F' ⟶ F`. --/ -@[simps] -def natTransEquiv {F F' : C ⥤ D} {G G' : D ⥤ C} (adj1 : F ⊣ G) (adj2 : F' ⊣ G') : - (G ⟶ G') ≃ (F' ⟶ F) where - toFun f := { - app := fun X ↦ F'.map ((adj1.unit ≫ whiskerLeft F f).app X) ≫ adj2.counit.app _ - naturality := by - intro X Y g - simp only [← Category.assoc, ← Functor.map_comp] - erw [(adj1.unit ≫ (whiskerLeft F f)).naturality] - simp - } - invFun f := { - app := fun X ↦ adj2.unit.app (G.obj X) ≫ G'.map (f.app (G.obj X) ≫ adj1.counit.app X) - naturality := by - intro X Y g - erw [← adj2.unit_naturality_assoc] - simp only [← Functor.map_comp] - simp - } - left_inv f := by - ext X - simp only [Functor.comp_obj, NatTrans.comp_app, Functor.id_obj, whiskerLeft_app, - Functor.map_comp, Category.assoc, unit_naturality_assoc, right_triangle_components_assoc] - erw [← f.naturality (adj1.counit.app X), ← Category.assoc] - simp - right_inv f := by - ext - simp - -@[simp] -lemma natTransEquiv_id {F : C ⥤ D} {G : D ⥤ C} (adj : F ⊣ G) : - natTransEquiv adj adj (𝟙 _) = 𝟙 _ := by ext; simp - -@[simp] -lemma natTransEquiv_id_symm {F : C ⥤ D} {G : D ⥤ C} (adj : F ⊣ G) : - (natTransEquiv adj adj).symm (𝟙 _) = 𝟙 _ := by ext; simp - -@[simp] -lemma natTransEquiv_comp {F F' F'' : C ⥤ D} {G G' G'' : D ⥤ C} - (adj1 : F ⊣ G) (adj2 : F' ⊣ G') (adj3 : F'' ⊣ G'') (f : G ⟶ G') (g : G' ⟶ G'') : - natTransEquiv adj2 adj3 g ≫ natTransEquiv adj1 adj2 f = natTransEquiv adj1 adj3 (f ≫ g) := by - apply (natTransEquiv adj1 adj3).symm.injective - ext X - simp only [natTransEquiv_symm_apply_app, Functor.comp_obj, NatTrans.comp_app, - natTransEquiv_apply_app, Functor.id_obj, whiskerLeft_app, Functor.map_comp, Category.assoc, - unit_naturality_assoc, right_triangle_components_assoc, Equiv.symm_apply_apply, - ← g.naturality_assoc, ← g.naturality] - simp only [← Category.assoc, unit_naturality, Functor.comp_obj, right_triangle_components, - Category.comp_id, ← f.naturality, Category.id_comp] - -@[simp] -lemma natTransEquiv_comp_symm {F F' F'' : C ⥤ D} {G G' G'' : D ⥤ C} - (adj1 : F ⊣ G) (adj2 : F' ⊣ G') (adj3 : F'' ⊣ G'') (f : F' ⟶ F) (g : F'' ⟶ F') : - (natTransEquiv adj1 adj2).symm f ≫ (natTransEquiv adj2 adj3).symm g = - (natTransEquiv adj1 adj3).symm (g ≫ f) := by - apply (natTransEquiv adj1 adj3).injective - ext - simp - -/-- -If `F ⊣ G` and `F' ⊣ G'` are adjunctions, then giving a natural isomorphism `G ≅ G'` is the -same as giving a natural transformation `F' ≅ F`. --/ -@[simps] -def natIsoEquiv {F F' : C ⥤ D} {G G' : D ⥤ C} (adj1 : F ⊣ G) (adj2 : F' ⊣ G') : - (G ≅ G') ≃ (F' ≅ F) where - toFun i := { - hom := natTransEquiv adj1 adj2 i.hom - inv := natTransEquiv adj2 adj1 i.inv - } - invFun i := { - hom := (natTransEquiv adj1 adj2).symm i.hom - inv := (natTransEquiv adj2 adj1).symm i.inv } - left_inv i := by simp - right_inv i := by simp - /-- If `F` and `F'` are both left adjoint to `G`, then they are naturally isomorphic. -/ def leftAdjointUniq {F F' : C ⥤ D} {G : D ⥤ C} (adj1 : F ⊣ G) (adj2 : F' ⊣ G) : F ≅ F' := - (natIsoEquiv adj1 adj2 (Iso.refl _)).symm + ((conjugateIsoEquiv adj1 adj2).symm (Iso.refl G)).symm -- Porting note (#10618): removed simp as simp can prove this theorem homEquiv_leftAdjointUniq_hom_app {F F' : C ⥤ D} {G : D ⥤ C} (adj1 : F ⊣ G) (adj2 : F' ⊣ G) @@ -141,9 +52,10 @@ theorem unit_leftAdjointUniq_hom_app theorem leftAdjointUniq_hom_counit {F F' : C ⥤ D} {G : D ⥤ C} (adj1 : F ⊣ G) (adj2 : F' ⊣ G) : whiskerLeft G (leftAdjointUniq adj1 adj2).hom ≫ adj2.counit = adj1.counit := by ext x - simp only [Functor.comp_obj, Functor.id_obj, leftAdjointUniq, Iso.symm_hom, natIsoEquiv_apply_inv, - Iso.refl_inv, NatTrans.comp_app, whiskerLeft_app, natTransEquiv_apply_app, whiskerLeft_id', - Category.comp_id, Category.assoc] + simp only [Functor.comp_obj, Functor.id_obj, leftAdjointUniq, Iso.symm_hom, + conjugateIsoEquiv_symm_apply_inv, Iso.refl_inv, NatTrans.comp_app, whiskerLeft_app, + conjugateEquiv_symm_apply_app, NatTrans.id_app, Functor.map_id, Category.id_comp, + Category.assoc] rw [← adj1.counit_naturality, ← Category.assoc, ← F.map_comp] simp @@ -180,7 +92,7 @@ theorem leftAdjointUniq_refl {F : C ⥤ D} {G : D ⥤ C} (adj1 : F ⊣ G) : /-- If `G` and `G'` are both right adjoint to `F`, then they are naturally isomorphic. -/ def rightAdjointUniq {F : C ⥤ D} {G G' : D ⥤ C} (adj1 : F ⊣ G) (adj2 : F ⊣ G') : G ≅ G' := - (natIsoEquiv adj1 adj2).symm (Iso.refl _) + conjugateIsoEquiv adj1 adj2 (Iso.refl _) -- Porting note (#10618): simp can prove this theorem homEquiv_symm_rightAdjointUniq_hom_app {F : C ⥤ D} {G G' : D ⥤ C} (adj1 : F ⊣ G) @@ -192,8 +104,8 @@ theorem homEquiv_symm_rightAdjointUniq_hom_app {F : C ⥤ D} {G G' : D ⥤ C} (a theorem unit_rightAdjointUniq_hom_app {F : C ⥤ D} {G G' : D ⥤ C} (adj1 : F ⊣ G) (adj2 : F ⊣ G') (x : C) : adj1.unit.app x ≫ (rightAdjointUniq adj1 adj2).hom.app (F.obj x) = adj2.unit.app x := by - simp only [Functor.id_obj, Functor.comp_obj, rightAdjointUniq, natIsoEquiv_symm_apply_hom, - Iso.refl_hom, natTransEquiv_symm_apply_app, NatTrans.id_app, Category.id_comp] + simp only [Functor.id_obj, Functor.comp_obj, rightAdjointUniq, conjugateIsoEquiv_apply_hom, + Iso.refl_hom, conjugateEquiv_apply_app, NatTrans.id_app, Functor.map_id, Category.id_comp] rw [← adj2.unit_naturality_assoc, ← G'.map_comp] simp @@ -243,4 +155,7 @@ theorem rightAdjointUniq_refl {F : C ⥤ D} {G : D ⥤ C} (adj1 : F ⊣ G) : end Adjunction +@[deprecated (since := "2024-10-07")] alias Adjunction.natTransEquiv := conjugateEquiv +@[deprecated (since := "2024-10-07")] alias Adjunction.natIsoEquiv := conjugateIsoEquiv + end CategoryTheory From 2611c37dc4c68bedce67f7dd4e66b44b03fba4d4 Mon Sep 17 00:00:00 2001 From: damiano Date: Mon, 7 Oct 2024 13:31:23 +0000 Subject: [PATCH 19/43] chore: remove more vars (#17489) --- Mathlib/Algebra/Group/Hom/CompTypeclasses.lean | 1 - Mathlib/Algebra/Order/Group/MinMax.lean | 4 ++-- Mathlib/Algebra/Order/Sub/Unbundled/Hom.lean | 4 ++-- Mathlib/Algebra/Ring/Equiv.lean | 6 +++--- Mathlib/CategoryTheory/Groupoid/FreeGroupoid.lean | 2 +- Mathlib/CategoryTheory/SmallObject/Iteration.lean | 4 ++-- Mathlib/Data/List/Basic.lean | 7 ++----- Mathlib/Data/List/Lattice.lean | 2 +- Mathlib/Data/List/Nodup.lean | 2 +- Mathlib/Order/Booleanisation.lean | 2 +- Mathlib/Order/Hom/Set.lean | 6 +++--- Mathlib/Order/Interval/Set/UnorderedInterval.lean | 8 ++++---- Mathlib/Tactic/CategoryTheory/Bicategory/Datatypes.lean | 2 +- 13 files changed, 23 insertions(+), 27 deletions(-) diff --git a/Mathlib/Algebra/Group/Hom/CompTypeclasses.lean b/Mathlib/Algebra/Group/Hom/CompTypeclasses.lean index b8c7e91b6a968..0bdd4b963dd4a 100644 --- a/Mathlib/Algebra/Group/Hom/CompTypeclasses.lean +++ b/Mathlib/Algebra/Group/Hom/CompTypeclasses.lean @@ -48,7 +48,6 @@ attribute [simp] CompTriple.comp_eq namespace CompTriple -variable {M' : Type*} [Monoid M'] variable {M N P : Type*} [Monoid M] [Monoid N] [Monoid P] /-- Class of Id maps -/ diff --git a/Mathlib/Algebra/Order/Group/MinMax.lean b/Mathlib/Algebra/Order/Group/MinMax.lean index 4e53be43c3ca4..adf6af1c4b904 100644 --- a/Mathlib/Algebra/Order/Group/MinMax.lean +++ b/Mathlib/Algebra/Order/Group/MinMax.lean @@ -30,7 +30,7 @@ end section LinearOrderedCommGroup -variable {α : Type*} [LinearOrderedCommGroup α] {a b c : α} +variable {α : Type*} [LinearOrderedCommGroup α] @[to_additive min_neg_neg] theorem min_inv_inv' (a b : α) : min a⁻¹ b⁻¹ = (max a b)⁻¹ := @@ -64,7 +64,7 @@ end LinearOrderedCommGroup section LinearOrderedAddCommGroup -variable {α : Type*} [LinearOrderedAddCommGroup α] {a b c : α} +variable {α : Type*} [LinearOrderedAddCommGroup α] theorem max_sub_max_le_max (a b c d : α) : max a b - max c d ≤ max (a - c) (b - d) := by simp only [sub_le_iff_le_add, max_le_iff]; constructor diff --git a/Mathlib/Algebra/Order/Sub/Unbundled/Hom.lean b/Mathlib/Algebra/Order/Sub/Unbundled/Hom.lean index 2bf9585fa56a6..fedfe17b0c39c 100644 --- a/Mathlib/Algebra/Order/Sub/Unbundled/Hom.lean +++ b/Mathlib/Algebra/Order/Sub/Unbundled/Hom.lean @@ -16,7 +16,7 @@ variable {α β : Type*} section Add -variable [Preorder α] [Add α] [Sub α] [OrderedSub α] {a b c d : α} +variable [Preorder α] [Add α] [Sub α] [OrderedSub α] theorem AddHom.le_map_tsub [Preorder β] [Add β] [Sub β] [OrderedSub β] (f : AddHom α β) (hf : Monotone f) (a b : α) : f a - f b ≤ f (a - b) := by @@ -49,7 +49,7 @@ theorem OrderIso.map_tsub {M N : Type*} [Preorder M] [Add M] [Sub M] [OrderedSub section Preorder variable [Preorder α] -variable [AddCommMonoid α] [Sub α] [OrderedSub α] {a b c d : α} +variable [AddCommMonoid α] [Sub α] [OrderedSub α] theorem AddMonoidHom.le_map_tsub [Preorder β] [AddCommMonoid β] [Sub β] [OrderedSub β] (f : α →+ β) (hf : Monotone f) (a b : α) : f a - f b ≤ f (a - b) := diff --git a/Mathlib/Algebra/Ring/Equiv.lean b/Mathlib/Algebra/Ring/Equiv.lean index eea3f38baa20c..ae83ce7dd57b1 100644 --- a/Mathlib/Algebra/Ring/Equiv.lean +++ b/Mathlib/Algebra/Ring/Equiv.lean @@ -409,7 +409,7 @@ end Opposite section NonUnitalSemiring -variable [NonUnitalNonAssocSemiring R] [NonUnitalNonAssocSemiring S] (f : R ≃+* S) (x y : R) +variable [NonUnitalNonAssocSemiring R] [NonUnitalNonAssocSemiring S] (f : R ≃+* S) (x : R) /-- A ring isomorphism sends zero to zero. -/ protected theorem map_zero : f 0 = 0 := @@ -533,7 +533,7 @@ end NonUnitalSemiring section Semiring -variable [NonAssocSemiring R] [NonAssocSemiring S] (f : R ≃+* S) (x y : R) +variable [NonAssocSemiring R] [NonAssocSemiring S] (f : R ≃+* S) (x : R) /-- A ring isomorphism sends one to one. -/ protected theorem map_one : f 1 = 1 := @@ -604,7 +604,7 @@ end NonUnitalRing section Ring -variable [NonAssocRing R] [NonAssocRing S] (f : R ≃+* S) (x y : R) +variable [NonAssocRing R] [NonAssocRing S] (f : R ≃+* S) -- Porting note (#10618): `simp` can now prove that, so we remove the `@[simp]` tag theorem map_neg_one : f (-1) = -1 := diff --git a/Mathlib/CategoryTheory/Groupoid/FreeGroupoid.lean b/Mathlib/CategoryTheory/Groupoid/FreeGroupoid.lean index 54aa24e65b239..79f597c240e2e 100644 --- a/Mathlib/CategoryTheory/Groupoid/FreeGroupoid.lean +++ b/Mathlib/CategoryTheory/Groupoid/FreeGroupoid.lean @@ -132,7 +132,7 @@ theorem of_eq : section UniversalProperty -variable {V' : Type u'} [Groupoid V'] (φ : V ⥤q V') +variable {V' : Type u'} [Groupoid V'] /-- The lift of a prefunctor to a groupoid, to a functor from `FreeGroupoid V` -/ def lift (φ : V ⥤q V') : FreeGroupoid V ⥤ V' := diff --git a/Mathlib/CategoryTheory/SmallObject/Iteration.lean b/Mathlib/CategoryTheory/SmallObject/Iteration.lean index 4e9947245dd1b..a9470418b418d 100644 --- a/Mathlib/CategoryTheory/SmallObject/Iteration.lean +++ b/Mathlib/CategoryTheory/SmallObject/Iteration.lean @@ -123,7 +123,7 @@ lemma mapSucc_eq (i : J) (hi : i < j) : end -variable (iter₁ iter₂ iter₃ : Φ.Iteration ε j) +variable (iter₁ iter₂ : Φ.Iteration ε j) /-- A morphism between two objects `iter₁` and `iter₂` in the category `Φ.Iteration ε j` of `j`th iterations of a functor `Φ` @@ -148,7 +148,7 @@ attribute [simp, reassoc] natTrans_app_zero def id : Hom iter₁ iter₁ where natTrans := 𝟙 _ -variable {iter₁ iter₂ iter₃} +variable {iter₁ iter₂} -- Note: this is not made a global ext lemma because it is shown below -- that the type of morphisms is a subsingleton. diff --git a/Mathlib/Data/List/Basic.lean b/Mathlib/Data/List/Basic.lean index 9c2f87605db92..43008b218c99b 100644 --- a/Mathlib/Data/List/Basic.lean +++ b/Mathlib/Data/List/Basic.lean @@ -1358,7 +1358,7 @@ theorem intersperse_cons_cons (a b c : α) (tl : List α) : section SplitAtOn -variable (p : α → Bool) (xs ys : List α) (ls : List (List α)) (f : List α → List α) +variable (p : α → Bool) (xs : List α) (ls : List (List α)) attribute [simp] splitAt_eq @@ -1535,8 +1535,6 @@ theorem sizeOf_lt_sizeOf_of_mem [SizeOf α] {x : α} {l : List α} (hx : x ∈ l section find? -variable {p : α → Bool} {l : List α} {a : α} - @[deprecated (since := "2024-05-05")] alias find?_mem := mem_of_find?_eq_some end find? @@ -2278,8 +2276,7 @@ theorem disjoint_reverse_right {l₁ l₂ : List α} : Disjoint l₁ l₂.revers end Disjoint section lookup - -variable {α β : Type*} [BEq α] [LawfulBEq α] +variable [BEq α] [LawfulBEq α] lemma lookup_graph (f : α → β) {a : α} {as : List α} (h : a ∈ as) : lookup a (as.map fun x => (x, f x)) = some (f a) := by diff --git a/Mathlib/Data/List/Lattice.lean b/Mathlib/Data/List/Lattice.lean index 75dc0c1e112bf..5a0ae92ccc6c0 100644 --- a/Mathlib/Data/List/Lattice.lean +++ b/Mathlib/Data/List/Lattice.lean @@ -29,7 +29,7 @@ open Nat namespace List -variable {α : Type*} {l l₁ l₂ : List α} {p : α → Prop} {a : α} +variable {α : Type*} {l₁ l₂ : List α} {p : α → Prop} {a : α} /-! ### `Disjoint` -/ diff --git a/Mathlib/Data/List/Nodup.lean b/Mathlib/Data/List/Nodup.lean index 9d86abe857e93..e8c9adf26282f 100644 --- a/Mathlib/Data/List/Nodup.lean +++ b/Mathlib/Data/List/Nodup.lean @@ -18,7 +18,7 @@ universe u v open Function -variable {α : Type u} {β : Type v} {l l₁ l₂ : List α} {r : α → α → Prop} {a b : α} +variable {α : Type u} {β : Type v} {l l₁ l₂ : List α} {r : α → α → Prop} {a : α} namespace List diff --git a/Mathlib/Order/Booleanisation.lean b/Mathlib/Order/Booleanisation.lean index dd3aa408d0c38..aa7f7370581b4 100644 --- a/Mathlib/Order/Booleanisation.lean +++ b/Mathlib/Order/Booleanisation.lean @@ -58,7 +58,7 @@ instance instCompl : HasCompl (Booleanisation α) where @[simp] lemma compl_lift (a : α) : (lift a)ᶜ = comp a := rfl @[simp] lemma compl_comp (a : α) : (comp a)ᶜ = lift a := rfl -variable [GeneralizedBooleanAlgebra α] {x y : Booleanisation α} {a b : α} +variable [GeneralizedBooleanAlgebra α] {a b : α} /-- The order on `Booleanisation α` is as follows: For `a b : α`, * `a ≤ b` iff `a ≤ b` in `α` diff --git a/Mathlib/Order/Hom/Set.lean b/Mathlib/Order/Hom/Set.lean index 96809de232fb1..382a2354e4f5d 100644 --- a/Mathlib/Order/Hom/Set.lean +++ b/Mathlib/Order/Hom/Set.lean @@ -16,13 +16,13 @@ import Mathlib.Order.WellFounded open OrderDual -variable {F α β γ δ : Type*} +variable {α β : Type*} namespace OrderIso section LE -variable [LE α] [LE β] [LE γ] +variable [LE α] [LE β] theorem range_eq (e : α ≃o β) : Set.range e = Set.univ := e.surjective.range_eq @@ -58,7 +58,7 @@ end LE open Set -variable [Preorder α] [Preorder β] [Preorder γ] +variable [Preorder α] /-- Order isomorphism between two equal sets. -/ def setCongr (s t : Set α) (h : s = t) : diff --git a/Mathlib/Order/Interval/Set/UnorderedInterval.lean b/Mathlib/Order/Interval/Set/UnorderedInterval.lean index 8c14071911f3d..ea241a7fbf3dc 100644 --- a/Mathlib/Order/Interval/Set/UnorderedInterval.lean +++ b/Mathlib/Order/Interval/Set/UnorderedInterval.lean @@ -44,7 +44,7 @@ namespace Set section Lattice -variable [Lattice α] [Lattice β] {a a₁ a₂ b b₁ b₂ c x : α} +variable [Lattice α] [Lattice β] {a a₁ a₂ b b₁ b₂ x : α} /-- `uIcc a b` is the set of elements lying between `a` and `b`, with `a` and `b` included. Note that we define it more generally in a lattice as `Set.Icc (a ⊓ b) (a ⊔ b)`. In a product type, @@ -134,7 +134,7 @@ open Interval section DistribLattice -variable [DistribLattice α] {a a₁ a₂ b b₁ b₂ c x : α} +variable [DistribLattice α] {a b c : α} lemma eq_of_mem_uIcc_of_mem_uIcc (ha : a ∈ [[b, c]]) (hb : b ∈ [[a, c]]) : a = b := eq_of_inf_eq_sup_eq (inf_congr_right ha.1 hb.1) <| sup_congr_right ha.2 hb.2 @@ -155,7 +155,7 @@ section LinearOrder variable [LinearOrder α] section Lattice -variable [Lattice β] {f : α → β} {s : Set α} {a b : α} +variable [Lattice β] {f : α → β} {a b : α} lemma _root_.MonotoneOn.mapsTo_uIcc (hf : MonotoneOn f (uIcc a b)) : MapsTo f (uIcc a b) (uIcc (f a) (f b)) := by @@ -187,7 +187,7 @@ lemma _root_.Antitone.image_uIcc_subset (hf : Antitone f) : f '' uIcc a b ⊆ uI end Lattice -variable [LinearOrder β] {f : α → β} {s : Set α} {a a₁ a₂ b b₁ b₂ c d x : α} +variable [LinearOrder β] {f : α → β} {s : Set α} {a a₁ a₂ b b₁ b₂ c : α} theorem Icc_min_max : Icc (min a b) (max a b) = [[a, b]] := rfl diff --git a/Mathlib/Tactic/CategoryTheory/Bicategory/Datatypes.lean b/Mathlib/Tactic/CategoryTheory/Bicategory/Datatypes.lean index e7c2dd37d4f41..42225947c25da 100644 --- a/Mathlib/Tactic/CategoryTheory/Bicategory/Datatypes.lean +++ b/Mathlib/Tactic/CategoryTheory/Bicategory/Datatypes.lean @@ -103,7 +103,7 @@ instance : MonadMor₁ BicategoryM where section universe w v u -variable {B : Type u} [Bicategory.{w, v} B] {a b c d e : B} +variable {B : Type u} [Bicategory.{w, v} B] {a b c : B} theorem structuralIso_inv {f g : a ⟶ b} (η : f ≅ g) : η.symm.hom = η.inv := by From d60a43683b51f127838d4015e0603c9f9f7b3682 Mon Sep 17 00:00:00 2001 From: damiano Date: Mon, 7 Oct 2024 13:31:24 +0000 Subject: [PATCH 20/43] fix: comment more cache (#17491) The previous fix was an improvement, but was not aggressive enough in commenting the failing step. This step used to use `cache` to download the cache for `Mathlib.Init`. If this was successful, then it would continue trying to download all of the available cache. Otherwise, it would start with building. The underlying reason is that if even the cache for `Mathlib.Init` is not there, there is not much of an available cache and there is no need to spend a couple of minutes realizing as much. However, changes to `lake build --no-buld` (possibly) have made this step always fail to download any cache, so we revert to the old behaviour of trying to get the cache no matter what. --- .github/build.in.yml | 2 +- .github/workflows/bors.yml | 2 +- .github/workflows/build.yml | 2 +- .github/workflows/build_fork.yml | 2 +- 4 files changed, 4 insertions(+), 4 deletions(-) diff --git a/.github/build.in.yml b/.github/build.in.yml index 1b0a8caaa4199..3facc8b60a92b 100644 --- a/.github/build.in.yml +++ b/.github/build.in.yml @@ -134,7 +134,7 @@ jobs: rm -rf .lake/build/lib/Mathlib/ # Fail quickly if the cache is completely cold, by checking for Mathlib.Init lake exe cache get #Mathlib.Init - lake build --no-build Mathlib.Init && lake exe cache get || echo "No cache for 'Mathlib.Init' available" + #lake build --no-build Mathlib.Init && lake exe cache get || echo "No cache for 'Mathlib.Init' available" - name: update {Mathlib, Tactic, Counterexamples, Archive}.lean id: mk_all diff --git a/.github/workflows/bors.yml b/.github/workflows/bors.yml index 82c9c075d5fd7..207aee8482616 100644 --- a/.github/workflows/bors.yml +++ b/.github/workflows/bors.yml @@ -144,7 +144,7 @@ jobs: rm -rf .lake/build/lib/Mathlib/ # Fail quickly if the cache is completely cold, by checking for Mathlib.Init lake exe cache get #Mathlib.Init - lake build --no-build Mathlib.Init && lake exe cache get || echo "No cache for 'Mathlib.Init' available" + #lake build --no-build Mathlib.Init && lake exe cache get || echo "No cache for 'Mathlib.Init' available" - name: update {Mathlib, Tactic, Counterexamples, Archive}.lean id: mk_all diff --git a/.github/workflows/build.yml b/.github/workflows/build.yml index 5ae06124e19ce..c47cd306ed4da 100644 --- a/.github/workflows/build.yml +++ b/.github/workflows/build.yml @@ -151,7 +151,7 @@ jobs: rm -rf .lake/build/lib/Mathlib/ # Fail quickly if the cache is completely cold, by checking for Mathlib.Init lake exe cache get #Mathlib.Init - lake build --no-build Mathlib.Init && lake exe cache get || echo "No cache for 'Mathlib.Init' available" + #lake build --no-build Mathlib.Init && lake exe cache get || echo "No cache for 'Mathlib.Init' available" - name: update {Mathlib, Tactic, Counterexamples, Archive}.lean id: mk_all diff --git a/.github/workflows/build_fork.yml b/.github/workflows/build_fork.yml index e3e86f31e2e1e..c431ad5e7d2f2 100644 --- a/.github/workflows/build_fork.yml +++ b/.github/workflows/build_fork.yml @@ -148,7 +148,7 @@ jobs: rm -rf .lake/build/lib/Mathlib/ # Fail quickly if the cache is completely cold, by checking for Mathlib.Init lake exe cache get #Mathlib.Init - lake build --no-build Mathlib.Init && lake exe cache get || echo "No cache for 'Mathlib.Init' available" + #lake build --no-build Mathlib.Init && lake exe cache get || echo "No cache for 'Mathlib.Init' available" - name: update {Mathlib, Tactic, Counterexamples, Archive}.lean id: mk_all From 5017a681823e948dea4ad55740d86cc8754e4ea6 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Mon, 7 Oct 2024 14:27:57 +0000 Subject: [PATCH 21/43] refactor(LinearAlgebra/Eigenspace): unified definition of (max(gen)?)?Eigenspace (#16025) --- Mathlib/Algebra/DirectSum/LinearMap.lean | 22 +- Mathlib/Algebra/Lie/TraceForm.lean | 10 +- Mathlib/Algebra/Lie/Weights/Basic.lean | 6 +- .../InnerProductSpace/JointEigenspace.lean | 2 +- Mathlib/LinearAlgebra/Eigenspace/Basic.lean | 629 ++++++++++++++---- Mathlib/LinearAlgebra/Eigenspace/Minpoly.lean | 2 +- .../Eigenspace/Triangularizable.lean | 4 +- Mathlib/Order/Hom/Basic.lean | 14 + Mathlib/RingTheory/Nilpotent/Lemmas.lean | 20 +- 9 files changed, 543 insertions(+), 166 deletions(-) diff --git a/Mathlib/Algebra/DirectSum/LinearMap.lean b/Mathlib/Algebra/DirectSum/LinearMap.lean index c3b5c570dd966..7ca4b191553b8 100644 --- a/Mathlib/Algebra/DirectSum/LinearMap.lean +++ b/Mathlib/Algebra/DirectSum/LinearMap.lean @@ -100,27 +100,27 @@ lemma trace_comp_eq_zero_of_commute_of_trace_restrict_eq_zero [IsDomain R] [IsPrincipalIdealRing R] [Module.Free R M] [Module.Finite R M] {f g : Module.End R M} (h_comm : Commute f g) - (hf : ⨆ μ, ⨆ k, f.genEigenspace μ k = ⊤) - (hg : ∀ μ, trace R _ (g.restrict (f.mapsTo_iSup_genEigenspace_of_comm h_comm μ)) = 0) : + (hf : ⨆ μ, f.maxGenEigenspace μ = ⊤) + (hg : ∀ μ, trace R _ (g.restrict (f.mapsTo_maxGenEigenspace_of_comm h_comm μ)) = 0) : trace R _ (g ∘ₗ f) = 0 := by have hfg : ∀ μ, - MapsTo (g ∘ₗ f) ↑(⨆ k, f.genEigenspace μ k) ↑(⨆ k, f.genEigenspace μ k) := - fun μ ↦ (f.mapsTo_iSup_genEigenspace_of_comm h_comm μ).comp - (f.mapsTo_iSup_genEigenspace_of_comm rfl μ) + MapsTo (g ∘ₗ f) ↑(f.maxGenEigenspace μ) ↑(f.maxGenEigenspace μ) := + fun μ ↦ (f.mapsTo_maxGenEigenspace_of_comm h_comm μ).comp + (f.mapsTo_maxGenEigenspace_of_comm rfl μ) suffices ∀ μ, trace R _ ((g ∘ₗ f).restrict (hfg μ)) = 0 by classical have hds := DirectSum.isInternal_submodule_of_independent_of_iSup_eq_top - f.independent_genEigenspace hf - have h_fin : {μ | ⨆ k, f.genEigenspace μ k ≠ ⊥}.Finite := + f.independent_maxGenEigenspace hf + have h_fin : {μ | f.maxGenEigenspace μ ≠ ⊥}.Finite := CompleteLattice.WellFounded.finite_ne_bot_of_independent IsWellFounded.wf - f.independent_genEigenspace + f.independent_maxGenEigenspace simp [trace_eq_sum_trace_restrict' hds h_fin hfg, this] intro μ - replace h_comm : Commute (g.restrict (f.mapsTo_iSup_genEigenspace_of_comm h_comm μ)) - (f.restrict (f.mapsTo_iSup_genEigenspace_of_comm rfl μ)) := + replace h_comm : Commute (g.restrict (f.mapsTo_maxGenEigenspace_of_comm h_comm μ)) + (f.restrict (f.mapsTo_maxGenEigenspace_of_comm rfl μ)) := restrict_commute h_comm.symm _ _ rw [restrict_comp, trace_comp_eq_mul_of_commute_of_isNilpotent μ h_comm - (f.isNilpotent_restrict_iSup_sub_algebraMap μ), hg, mul_zero] + (f.isNilpotent_restrict_maxGenEigenspace_sub_algebraMap μ), hg, mul_zero] lemma mapsTo_biSup_of_mapsTo {ι : Type*} {N : ι → Submodule R M} (s : Set ι) {f : Module.End R M} (hf : ∀ i, MapsTo f (N i) (N i)) : diff --git a/Mathlib/Algebra/Lie/TraceForm.lean b/Mathlib/Algebra/Lie/TraceForm.lean index 6344226a618c7..bddebc215744c 100644 --- a/Mathlib/Algebra/Lie/TraceForm.lean +++ b/Mathlib/Algebra/Lie/TraceForm.lean @@ -226,9 +226,9 @@ lemma traceForm_eq_sum_genWeightSpaceOf convert finite_genWeightSpaceOf_ne_bot R L M z exact LieSubmodule.coeSubmodule_eq_bot_iff (genWeightSpaceOf M _ _) classical - have hds := DirectSum.isInternal_submodule_of_independent_of_iSup_eq_top - (LieSubmodule.independent_iff_coe_toSubmodule.mp <| independent_genWeightSpaceOf R L M z) - (IsTriangularizable.iSup_eq_top z) + have h := LieSubmodule.independent_iff_coe_toSubmodule.mp <| independent_genWeightSpaceOf R L M z + have hds := DirectSum.isInternal_submodule_of_independent_of_iSup_eq_top h <| by + simp [← LieSubmodule.iSup_coe_toSubmodule] simp only [LinearMap.coeFn_sum, Finset.sum_apply, traceForm_apply_apply, LinearMap.trace_eq_sum_trace_restrict' hds hfin hxy] exact Finset.sum_congr (by simp) (fun χ _ ↦ rfl) @@ -276,9 +276,9 @@ lemma lowerCentralSeries_one_inf_center_le_ker_traceForm [Module.Free R M] [Modu intro y exact y.induction_on rfl (fun a u ↦ by simp [hzc u]) (fun u v hu hv ↦ by simp [hu, hv]) apply LinearMap.trace_comp_eq_zero_of_commute_of_trace_restrict_eq_zero - · exact IsTriangularizable.iSup_eq_top (1 ⊗ₜ[R] x) + · simpa only [Module.End.maxGenEigenspace_def] using IsTriangularizable.iSup_eq_top (1 ⊗ₜ[R] x) · exact fun μ ↦ trace_toEnd_eq_zero_of_mem_lcs A (A ⊗[R] L) - (genWeightSpaceOf (A ⊗[R] M) μ (1 ⊗ₜ x)) (le_refl 1) hz + (genWeightSpaceOf (A ⊗[R] M) μ ((1:A) ⊗ₜ[R] x)) (le_refl 1) hz · exact commute_toEnd_of_mem_center_right (A ⊗[R] M) hzc (1 ⊗ₜ x) /-- A nilpotent Lie algebra with a representation whose trace form is non-singular is Abelian. -/ diff --git a/Mathlib/Algebra/Lie/Weights/Basic.lean b/Mathlib/Algebra/Lie/Weights/Basic.lean index d991a7a0c5937..090ee19929b33 100644 --- a/Mathlib/Algebra/Lie/Weights/Basic.lean +++ b/Mathlib/Algebra/Lie/Weights/Basic.lean @@ -782,8 +782,10 @@ lemma iSup_genWeightSpace_eq_top [IsTriangularizable K L M] : ⨆ χ : L → K, genWeightSpace M χ = ⊤ := by simp only [← LieSubmodule.coe_toSubmodule_eq_iff, LieSubmodule.iSup_coe_toSubmodule, LieSubmodule.iInf_coe_toSubmodule, LieSubmodule.top_coeSubmodule, genWeightSpace] - exact Module.End.iSup_iInf_maxGenEigenspace_eq_top_of_forall_mapsTo (toEnd K L M) - (fun x y φ z ↦ (genWeightSpaceOf M φ y).lie_mem) (IsTriangularizable.iSup_eq_top) + refine Module.End.iSup_iInf_maxGenEigenspace_eq_top_of_forall_mapsTo (toEnd K L M) + (fun x y φ z ↦ (genWeightSpaceOf M φ y).lie_mem) ?_ + simp_rw [Module.End.maxGenEigenspace_def] + apply IsTriangularizable.iSup_eq_top lemma iSup_genWeightSpace_eq_top' [IsTriangularizable K L M] : ⨆ χ : Weight K L M, genWeightSpace M χ = ⊤ := by diff --git a/Mathlib/Analysis/InnerProductSpace/JointEigenspace.lean b/Mathlib/Analysis/InnerProductSpace/JointEigenspace.lean index 868496a0b3462..049ba1532a238 100644 --- a/Mathlib/Analysis/InnerProductSpace/JointEigenspace.lean +++ b/Mathlib/Analysis/InnerProductSpace/JointEigenspace.lean @@ -49,7 +49,7 @@ variable {α : 𝕜} {A B : E →ₗ[𝕜] E} theorem eigenspace_invariant_of_commute (hAB : A ∘ₗ B = B ∘ₗ A) (α : 𝕜) : ∀ v ∈ (eigenspace A α), (B v ∈ eigenspace A α) := by intro v hv - rw [eigenspace, mem_ker, sub_apply, Module.algebraMap_end_apply, ← comp_apply A B v, hAB, + rw [eigenspace_def, mem_ker, sub_apply, smul_apply, one_apply, ← comp_apply A B v, hAB, comp_apply B A v, ← map_smul, ← map_sub, hv, map_zero] at * /--The simultaneous eigenspaces of a pair of commuting symmetric operators form an diff --git a/Mathlib/LinearAlgebra/Eigenspace/Basic.lean b/Mathlib/LinearAlgebra/Eigenspace/Basic.lean index 3effed50341ef..e2e7aa9fcb4c9 100644 --- a/Mathlib/LinearAlgebra/Eigenspace/Basic.lean +++ b/Mathlib/LinearAlgebra/Eigenspace/Basic.lean @@ -58,227 +58,525 @@ open Module Set variable {K R : Type v} {V M : Type w} [CommRing R] [AddCommGroup M] [Module R M] [Field K] [AddCommGroup V] [Module K V] +/-- The submodule `unifEigenspace f μ k` for a linear map `f`, a scalar `μ`, +and a number `k : ℕ∞` is the kernel of `(f - μ • id) ^ k` if `k` is a natural number, +or the union of all these kernels if `k = ∞`. -/ +def unifEigenspace (f : End R M) (μ : R) : ℕ∞ →o Submodule R M where + toFun k := ⨆ l : ℕ, ⨆ _ : l ≤ k, LinearMap.ker ((f - μ • 1) ^ l) + monotone' _ _ hkl := biSup_mono fun _ hi ↦ hi.trans hkl + +lemma mem_unifEigenspace {f : End R M} {μ : R} {k : ℕ∞} {x : M} : + x ∈ f.unifEigenspace μ k ↔ ∃ l : ℕ, l ≤ k ∧ x ∈ LinearMap.ker ((f - μ • 1) ^ l) := by + have : Nonempty {l : ℕ // l ≤ k} := ⟨⟨0, zero_le _⟩⟩ + have : Directed (ι := { i : ℕ // i ≤ k }) (· ≤ ·) fun i ↦ LinearMap.ker ((f - μ • 1) ^ (i : ℕ)) := + Monotone.directed_le fun m n h ↦ by simpa using (f - μ • 1).iterateKer.monotone h + simp_rw [unifEigenspace, OrderHom.coe_mk, LinearMap.mem_ker, iSup_subtype', + Submodule.mem_iSup_of_directed _ this, LinearMap.mem_ker, Subtype.exists, exists_prop] + +lemma unifEigenspace_directed {f : End R M} {μ : R} {k : ℕ∞} : + Directed (· ≤ ·) (fun l : {l : ℕ // l ≤ k} ↦ f.unifEigenspace μ l) := by + have aux : Monotone ((↑) : {l : ℕ // l ≤ k} → ℕ∞) := fun x y h ↦ by simpa using h + exact ((unifEigenspace f μ).monotone.comp aux).directed_le + +lemma mem_unifEigenspace_nat {f : End R M} {μ : R} {k : ℕ} {x : M} : + x ∈ f.unifEigenspace μ k ↔ x ∈ LinearMap.ker ((f - μ • 1) ^ k) := by + rw [mem_unifEigenspace] + constructor + · rintro ⟨l, hl, hx⟩ + simp only [Nat.cast_le] at hl + exact (f - μ • 1).iterateKer.monotone hl hx + · intro hx + exact ⟨k, le_rfl, hx⟩ + +lemma mem_unifEigenspace_top {f : End R M} {μ : R} {x : M} : + x ∈ f.unifEigenspace μ ⊤ ↔ ∃ k : ℕ, x ∈ LinearMap.ker ((f - μ • 1) ^ k) := by + simp [mem_unifEigenspace] + +lemma unifEigenspace_nat {f : End R M} {μ : R} {k : ℕ} : + f.unifEigenspace μ k = LinearMap.ker ((f - μ • 1) ^ k) := by + ext; simp [mem_unifEigenspace_nat] + +lemma unifEigenspace_eq_iSup_unifEigenspace_nat (f : End R M) (μ : R) (k : ℕ∞) : + f.unifEigenspace μ k = ⨆ l : {l : ℕ // l ≤ k}, f.unifEigenspace μ l := by + simp_rw [unifEigenspace_nat, unifEigenspace, OrderHom.coe_mk, iSup_subtype] + +lemma unifEigenspace_top (f : End R M) (μ : R) : + f.unifEigenspace μ ⊤ = ⨆ k : ℕ, f.unifEigenspace μ k := by + rw [unifEigenspace_eq_iSup_unifEigenspace_nat, iSup_subtype] + simp only [le_top, iSup_pos, OrderHom.coe_mk] + +lemma unifEigenspace_one {f : End R M} {μ : R} : + f.unifEigenspace μ 1 = LinearMap.ker (f - μ • 1) := by + rw [← Nat.cast_one, unifEigenspace_nat, pow_one] + +@[simp] +lemma mem_unifEigenspace_one {f : End R M} {μ : R} {x : M} : + x ∈ f.unifEigenspace μ 1 ↔ f x = μ • x := by + rw [unifEigenspace_one, LinearMap.mem_ker, LinearMap.sub_apply, + sub_eq_zero, LinearMap.smul_apply, LinearMap.one_apply] + +-- `simp` can prove this using `unifEigenspace_zero` +lemma mem_unifEigenspace_zero {f : End R M} {μ : R} {x : M} : + x ∈ f.unifEigenspace μ 0 ↔ x = 0 := by + rw [← Nat.cast_zero, mem_unifEigenspace_nat, pow_zero, LinearMap.mem_ker, LinearMap.one_apply] + +@[simp] +lemma unifEigenspace_zero {f : End R M} {μ : R} : + f.unifEigenspace μ 0 = ⊥ := by + ext; apply mem_unifEigenspace_zero + +@[simp] +lemma unifEigenspace_zero_nat (f : End R M) (k : ℕ) : + f.unifEigenspace 0 k = LinearMap.ker (f ^ k) := by + ext; simp [mem_unifEigenspace_nat] + +/-- Let `M` be an `R`-module, and `f` an `R`-linear endomorphism of `M`, +and let `μ : R` and `k : ℕ∞` be given. +Then `x : M` satisfies `HasUnifEigenvector f μ k x` if +`x ∈ f.unifEigenspace μ k` and `x ≠ 0`. + +For `k = 1`, this means that `x` is an eigenvector of `f` with eigenvalue `μ`. -/ +def HasUnifEigenvector (f : End R M) (μ : R) (k : ℕ∞) (x : M) : Prop := + x ∈ f.unifEigenspace μ k ∧ x ≠ 0 + +/-- Let `M` be an `R`-module, and `f` an `R`-linear endomorphism of `M`. +Then `μ : R` and `k : ℕ∞` satisfy `HasUnifEigenvalue f μ k` if +`f.unifEigenspace μ k ≠ ⊥`. + +For `k = 1`, this means that `μ` is an eigenvalue of `f`. -/ +def HasUnifEigenvalue (f : End R M) (μ : R) (k : ℕ∞) : Prop := + f.unifEigenspace μ k ≠ ⊥ + +/-- Let `M` be an `R`-module, and `f` an `R`-linear endomorphism of `M`. +For `k : ℕ∞`, we define `UnifEigenvalues f k` to be the type of all +`μ : R` that satisfy `f.HasUnifEigenvalue μ k`. + +For `k = 1` this is the type of all eigenvalues of `f`. -/ +def UnifEigenvalues (f : End R M) (k : ℕ∞) : Type _ := + { μ : R // f.HasUnifEigenvalue μ k } + +/-- The underlying value of a bundled eigenvalue. -/ +@[coe] +def UnifEigenvalues.val (f : Module.End R M) (k : ℕ∞) : UnifEigenvalues f k → R := Subtype.val + +instance UnifEigenvalues.instCoeOut {f : Module.End R M} (k : ℕ∞) : + CoeOut (UnifEigenvalues f k) R where + coe := UnifEigenvalues.val f k + +instance UnivEigenvalues.instDecidableEq [DecidableEq R] (f : Module.End R M) (k : ℕ∞) : + DecidableEq (UnifEigenvalues f k) := + inferInstanceAs (DecidableEq (Subtype (fun x : R ↦ f.HasUnifEigenvalue x k))) + +lemma HasUnifEigenvector.hasUnifEigenvalue {f : End R M} {μ : R} {k : ℕ∞} {x : M} + (h : f.HasUnifEigenvector μ k x) : f.HasUnifEigenvalue μ k := by + rw [HasUnifEigenvalue, Submodule.ne_bot_iff] + use x; exact h + +lemma HasUnifEigenvector.apply_eq_smul {f : End R M} {μ : R} {x : M} + (hx : f.HasUnifEigenvector μ 1 x) : f x = μ • x := + mem_unifEigenspace_one.mp hx.1 + +lemma HasUnifEigenvector.pow_apply {f : End R M} {μ : R} {v : M} (hv : f.HasUnifEigenvector μ 1 v) + (n : ℕ) : (f ^ n) v = μ ^ n • v := by + induction n <;> simp [*, pow_succ f, hv.apply_eq_smul, smul_smul, pow_succ' μ] + +theorem HasUnifEigenvalue.exists_hasUnifEigenvector + {f : End R M} {μ : R} {k : ℕ∞} (hμ : f.HasUnifEigenvalue μ k) : + ∃ v, f.HasUnifEigenvector μ k v := + Submodule.exists_mem_ne_zero_of_ne_bot hμ + +lemma HasUnifEigenvalue.pow {f : End R M} {μ : R} (h : f.HasUnifEigenvalue μ 1) (n : ℕ) : + (f ^ n).HasUnifEigenvalue (μ ^ n) 1 := by + rw [HasUnifEigenvalue, Submodule.ne_bot_iff] + obtain ⟨m : M, hm⟩ := h.exists_hasUnifEigenvector + exact ⟨m, by simpa [mem_unifEigenspace_one] using hm.pow_apply n, hm.2⟩ + +/-- A nilpotent endomorphism has nilpotent eigenvalues. + +See also `LinearMap.isNilpotent_trace_of_isNilpotent`. -/ +lemma HasUnifEigenvalue.isNilpotent_of_isNilpotent [NoZeroSMulDivisors R M] {f : End R M} + (hfn : IsNilpotent f) {μ : R} (hf : f.HasUnifEigenvalue μ 1) : + IsNilpotent μ := by + obtain ⟨m : M, hm⟩ := hf.exists_hasUnifEigenvector + obtain ⟨n : ℕ, hn : f ^ n = 0⟩ := hfn + exact ⟨n, by simpa [hn, hm.2, eq_comm (a := (0 : M))] using hm.pow_apply n⟩ + +lemma HasUnifEigenvalue.mem_spectrum {f : End R M} {μ : R} (hμ : HasUnifEigenvalue f μ 1) : + μ ∈ spectrum R f := by + refine spectrum.mem_iff.mpr fun h_unit ↦ ?_ + set f' := LinearMap.GeneralLinearGroup.toLinearEquiv h_unit.unit + rcases hμ.exists_hasUnifEigenvector with ⟨v, hv⟩ + refine hv.2 ((LinearMap.ker_eq_bot'.mp f'.ker) v (?_ : μ • v - f v = 0)) + rw [hv.apply_eq_smul, sub_self] + +lemma hasUnifEigenvalue_iff_mem_spectrum [FiniteDimensional K V] {f : End K V} {μ : K} : + f.HasUnifEigenvalue μ 1 ↔ μ ∈ spectrum K f := by + rw [spectrum.mem_iff, IsUnit.sub_iff, LinearMap.isUnit_iff_ker_eq_bot, + HasUnifEigenvalue, unifEigenspace_one, ne_eq, not_iff_not] + simp [Submodule.ext_iff, LinearMap.mem_ker] + +alias ⟨_, HasUnifEigenvalue.of_mem_spectrum⟩ := hasUnifEigenvalue_iff_mem_spectrum + +lemma unifEigenspace_div (f : End K V) (a b : K) (hb : b ≠ 0) : + unifEigenspace f (a / b) 1 = LinearMap.ker (b • f - a • 1) := + calc + unifEigenspace f (a / b) 1 = unifEigenspace f (b⁻¹ * a) 1 := by rw [div_eq_mul_inv, mul_comm] + _ = LinearMap.ker (f - (b⁻¹ * a) • 1) := by rw [unifEigenspace_one] + _ = LinearMap.ker (f - b⁻¹ • a • 1) := by rw [smul_smul] + _ = LinearMap.ker (b • (f - b⁻¹ • a • 1)) := by rw [LinearMap.ker_smul _ b hb] + _ = LinearMap.ker (b • f - a • 1) := by rw [smul_sub, smul_inv_smul₀ hb] + +/-- The generalized eigenrange for a linear map `f`, a scalar `μ`, and an exponent `k ∈ ℕ∞` +is the range of `(f - μ • id) ^ k` if `k` is a natural number, +or the infimum of these ranges if `k = ∞`. -/ +def unifEigenrange (f : End R M) (μ : R) (k : ℕ∞) : Submodule R M := + ⨅ l : ℕ, ⨅ (_ : l ≤ k), LinearMap.range ((f - μ • 1) ^ l) + +lemma unifEigenrange_nat {f : End R M} {μ : R} {k : ℕ} : + f.unifEigenrange μ k = LinearMap.range ((f - μ • 1) ^ k) := by + ext x + simp only [unifEigenrange, Nat.cast_le, Submodule.mem_iInf, LinearMap.mem_range] + constructor + · intro h + exact h _ le_rfl + · rintro ⟨x, rfl⟩ i hi + have : k = i + (k - i) := by omega + rw [this, pow_add] + exact ⟨_, rfl⟩ + +/-- The exponent of a generalized eigenvalue is never 0. -/ +lemma HasUnifEigenvalue.exp_ne_zero {f : End R M} {μ : R} {k : ℕ} + (h : f.HasUnifEigenvalue μ k) : k ≠ 0 := by + rintro rfl + simp [HasUnifEigenvalue, Nat.cast_zero, unifEigenspace_zero] at h + +/-- If there exists a natural number `k` such that the kernel of `(f - μ • id) ^ k` is the +maximal generalized eigenspace, then this value is the least such `k`. If not, this value is not +meaningful. -/ +noncomputable def maxUnifEigenspaceIndex (f : End R M) (μ : R) := + monotonicSequenceLimitIndex <| (f.unifEigenspace μ).comp <| WithTop.coeOrderHom.toOrderHom + +/-- For an endomorphism of a Noetherian module, the maximal eigenspace is always of the form kernel +`(f - μ • id) ^ k` for some `k`. -/ +lemma unifEigenspace_top_eq_maxUnifEigenspaceIndex [h : IsNoetherian R M] (f : End R M) (μ : R) : + unifEigenspace f μ ⊤ = f.unifEigenspace μ (maxUnifEigenspaceIndex f μ) := by + rw [isNoetherian_iff] at h + have := WellFounded.iSup_eq_monotonicSequenceLimit h <| + (f.unifEigenspace μ).comp <| WithTop.coeOrderHom.toOrderHom + convert this using 1 + simp only [unifEigenspace, OrderHom.coe_mk, le_top, iSup_pos, OrderHom.comp_coe, + Function.comp_def] + rw [iSup_prod', iSup_subtype', ← sSup_range, ← sSup_range] + congr + aesop + +lemma unifEigenspace_le_unifEigenspace_maxUnifEigenspaceIndex [IsNoetherian R M] (f : End R M) + (μ : R) (k : ℕ∞) : + f.unifEigenspace μ k ≤ f.unifEigenspace μ (maxUnifEigenspaceIndex f μ) := by + rw [← unifEigenspace_top_eq_maxUnifEigenspaceIndex] + exact (f.unifEigenspace μ).monotone le_top + +/-- Generalized eigenspaces for exponents at least `finrank K V` are equal to each other. -/ +theorem unifEigenspace_eq_unifEigenspace_maxUnifEigenspaceIndex_of_le [IsNoetherian R M] + (f : End R M) (μ : R) {k : ℕ} (hk : maxUnifEigenspaceIndex f μ ≤ k) : + f.unifEigenspace μ k = f.unifEigenspace μ (maxUnifEigenspaceIndex f μ) := + le_antisymm + (unifEigenspace_le_unifEigenspace_maxUnifEigenspaceIndex _ _ _) + ((f.unifEigenspace μ).monotone <| by simpa using hk) + +/-- A generalized eigenvalue for some exponent `k` is also + a generalized eigenvalue for exponents larger than `k`. -/ +lemma HasUnifEigenvalue.le {f : End R M} {μ : R} {k m : ℕ∞} + (hm : k ≤ m) (hk : f.HasUnifEigenvalue μ k) : + f.HasUnifEigenvalue μ m := by + unfold HasUnifEigenvalue at * + contrapose! hk + rw [← le_bot_iff, ← hk] + exact (f.unifEigenspace _).monotone hm + +/-- A generalized eigenvalue for some exponent `k` is also + a generalized eigenvalue for positive exponents. -/ +lemma HasUnifEigenvalue.lt {f : End R M} {μ : R} {k m : ℕ∞} + (hm : 0 < m) (hk : f.HasUnifEigenvalue μ k) : + f.HasUnifEigenvalue μ m := by + apply HasUnifEigenvalue.le (k := 1) (Order.one_le_iff_pos.mpr hm) + intro contra; apply hk + rw [unifEigenspace_one, LinearMap.ker_eq_bot] at contra + rw [eq_bot_iff] + intro x hx + rw [mem_unifEigenspace] at hx + rcases hx with ⟨l, -, hx⟩ + rwa [LinearMap.ker_eq_bot.mpr] at hx + rw [LinearMap.coe_pow (f - μ • 1) l] + exact Function.Injective.iterate contra l + +/-- Generalized eigenvalues are actually just eigenvalues. -/ +@[simp] +lemma hasUnifEigenvalue_iff_hasUnifEigenvalue_one {f : End R M} {μ : R} {k : ℕ∞} (hk : 0 < k) : + f.HasUnifEigenvalue μ k ↔ f.HasUnifEigenvalue μ 1 := + ⟨HasUnifEigenvalue.lt zero_lt_one, HasUnifEigenvalue.lt hk⟩ + +lemma maxUnifEigenspaceIndex_le_finrank [FiniteDimensional K V] (f : End K V) (μ : K) : + maxUnifEigenspaceIndex f μ ≤ finrank K V := by + apply Nat.sInf_le + intro n hn + apply le_antisymm + · exact (f.unifEigenspace μ).monotone <| WithTop.coeOrderHom.monotone hn + · show (f.unifEigenspace μ) n ≤ (f.unifEigenspace μ) (finrank K V) + rw [unifEigenspace_nat, unifEigenspace_nat] + apply ker_pow_le_ker_pow_finrank + +/-- Every generalized eigenvector is a generalized eigenvector for exponent `finrank K V`. + (Lemma 8.11 of [axler2015]) -/ +lemma unifEigenspace_le_unifEigenspace_finrank [FiniteDimensional K V] (f : End K V) + (μ : K) (k : ℕ∞) : f.unifEigenspace μ k ≤ f.unifEigenspace μ (finrank K V) := by + calc f.unifEigenspace μ k + ≤ f.unifEigenspace μ ⊤ := (f.unifEigenspace _).monotone le_top + _ ≤ f.unifEigenspace μ (finrank K V) := by + rw [unifEigenspace_top_eq_maxUnifEigenspaceIndex] + exact (f.unifEigenspace _).monotone <| by simpa using maxUnifEigenspaceIndex_le_finrank f μ + +/-- Generalized eigenspaces for exponents at least `finrank K V` are equal to each other. -/ +theorem unifEigenspace_eq_unifEigenspace_finrank_of_le [FiniteDimensional K V] + (f : End K V) (μ : K) {k : ℕ} (hk : finrank K V ≤ k) : + f.unifEigenspace μ k = f.unifEigenspace μ (finrank K V) := + le_antisymm + (unifEigenspace_le_unifEigenspace_finrank _ _ _) + ((f.unifEigenspace μ).monotone <| by simpa using hk) + +lemma mapsTo_unifEigenspace_of_comm {f g : End R M} (h : Commute f g) (μ : R) (k : ℕ∞) : + MapsTo g (f.unifEigenspace μ k) (f.unifEigenspace μ k) := by + intro x hx + simp only [SetLike.mem_coe, mem_unifEigenspace, LinearMap.mem_ker] at hx ⊢ + rcases hx with ⟨l, hl, hx⟩ + replace h : Commute ((f - μ • (1 : End R M)) ^ l) g := + (h.sub_left <| Algebra.commute_algebraMap_left μ g).pow_left l + use l, hl + rw [← LinearMap.comp_apply, ← LinearMap.mul_eq_comp, h.eq, LinearMap.mul_eq_comp, + LinearMap.comp_apply, hx, map_zero] + +/-- The restriction of `f - μ • 1` to the `k`-fold generalized `μ`-eigenspace is nilpotent. -/ +lemma isNilpotent_restrict_unifEigenspace_nat (f : End R M) (μ : R) (k : ℕ) + (h : MapsTo (f - μ • (1 : End R M)) + (f.unifEigenspace μ k) (f.unifEigenspace μ k) := + mapsTo_unifEigenspace_of_comm (Algebra.mul_sub_algebraMap_commutes f μ) μ k) : + IsNilpotent ((f - μ • 1).restrict h) := by + use k + ext ⟨x, hx⟩ + rw [mem_unifEigenspace_nat] at hx + rw [LinearMap.zero_apply, ZeroMemClass.coe_zero, ZeroMemClass.coe_eq_zero, + LinearMap.pow_restrict, LinearMap.restrict_apply] + ext + simpa + +/-- The restriction of `f - μ • 1` to the generalized `μ`-eigenspace is nilpotent. -/ +lemma isNilpotent_restrict_unifEigenspace_top [IsNoetherian R M] (f : End R M) (μ : R) + (h : MapsTo (f - μ • (1 : End R M)) + (f.unifEigenspace μ ⊤) (f.unifEigenspace μ ⊤) := + mapsTo_unifEigenspace_of_comm (Algebra.mul_sub_algebraMap_commutes f μ) μ _) : + IsNilpotent ((f - μ • 1).restrict h) := by + apply isNilpotent_restrict_of_le + swap; apply isNilpotent_restrict_unifEigenspace_nat f μ (maxUnifEigenspaceIndex f μ) + rw [unifEigenspace_top_eq_maxUnifEigenspaceIndex] + /-- The submodule `eigenspace f μ` for a linear map `f` and a scalar `μ` consists of all vectors `x` such that `f x = μ • x`. (Def 5.36 of [axler2015])-/ -def eigenspace (f : End R M) (μ : R) : Submodule R M := - LinearMap.ker (f - algebraMap R (End R M) μ) +abbrev eigenspace (f : End R M) (μ : R) : Submodule R M := + f.unifEigenspace μ 1 -lemma eigenspace_def (f : End R M) (μ : R) : - f.eigenspace μ = LinearMap.ker (f - algebraMap R (End R M) μ) := rfl +lemma eigenspace_def {f : End R M} {μ : R} : + f.eigenspace μ = LinearMap.ker (f - μ • 1) := by + rw [eigenspace, unifEigenspace_one] @[simp] -theorem eigenspace_zero (f : End R M) : f.eigenspace 0 = LinearMap.ker f := by simp [eigenspace] +theorem eigenspace_zero (f : End R M) : f.eigenspace 0 = LinearMap.ker f := by + simp only [eigenspace, ← Nat.cast_one (R := ℕ∞), unifEigenspace_zero_nat, pow_one] /-- A nonzero element of an eigenspace is an eigenvector. (Def 5.7 of [axler2015]) -/ -def HasEigenvector (f : End R M) (μ : R) (x : M) : Prop := - x ∈ eigenspace f μ ∧ x ≠ 0 +abbrev HasEigenvector (f : End R M) (μ : R) (x : M) : Prop := + HasUnifEigenvector f μ 1 x lemma hasEigenvector_iff {f : End R M} {μ : R} {x : M} : - f.HasEigenvector μ x ↔ x ∈ eigenspace f μ ∧ x ≠ 0 := Iff.rfl + f.HasEigenvector μ x ↔ x ∈ f.eigenspace μ ∧ x ≠ 0 := Iff.rfl /-- A scalar `μ` is an eigenvalue for a linear map `f` if there are nonzero vectors `x` such that `f x = μ • x`. (Def 5.5 of [axler2015]) -/ -def HasEigenvalue (f : End R M) (a : R) : Prop := - eigenspace f a ≠ ⊥ +abbrev HasEigenvalue (f : End R M) (a : R) : Prop := + HasUnifEigenvalue f a 1 -lemma hasEigenvalue_iff (f : End R M) (μ : R) : f.HasEigenvalue μ ↔ eigenspace f μ ≠ ⊥ := Iff.rfl +lemma hasEigenvalue_iff {f : End R M} {μ : R} : + f.HasEigenvalue μ ↔ f.eigenspace μ ≠ ⊥ := Iff.rfl /-- The eigenvalues of the endomorphism `f`, as a subtype of `R`. -/ -def Eigenvalues (f : End R M) : Type _ := - { μ : R // f.HasEigenvalue μ } +abbrev Eigenvalues (f : End R M) : Type _ := + UnifEigenvalues f 1 @[coe] -def Eigenvalues.val (f : Module.End R M) : Eigenvalues f → R := Subtype.val - -instance Eigenvalues.instCoeOut {f : Module.End R M} : CoeOut (Eigenvalues f) R where - coe := Eigenvalues.val f - -instance Eigenvalues.instDecidableEq [DecidableEq R] (f : Module.End R M) : - DecidableEq (Eigenvalues f) := - inferInstanceAs (DecidableEq (Subtype (fun x : R => HasEigenvalue f x))) +abbrev Eigenvalues.val (f : Module.End R M) : Eigenvalues f → R := UnifEigenvalues.val f 1 theorem hasEigenvalue_of_hasEigenvector {f : End R M} {μ : R} {x : M} (h : HasEigenvector f μ x) : - HasEigenvalue f μ := by - rw [HasEigenvalue, Submodule.ne_bot_iff] - use x; exact h + HasEigenvalue f μ := + h.hasUnifEigenvalue -theorem mem_eigenspace_iff {f : End R M} {μ : R} {x : M} : x ∈ eigenspace f μ ↔ f x = μ • x := by - rw [eigenspace, LinearMap.mem_ker, LinearMap.sub_apply, algebraMap_end_apply, sub_eq_zero] +theorem mem_eigenspace_iff {f : End R M} {μ : R} {x : M} : x ∈ eigenspace f μ ↔ f x = μ • x := + mem_unifEigenspace_one +nonrec theorem HasEigenvector.apply_eq_smul {f : End R M} {μ : R} {x : M} (hx : f.HasEigenvector μ x) : f x = μ • x := - mem_eigenspace_iff.mp hx.1 + hx.apply_eq_smul +nonrec theorem HasEigenvector.pow_apply {f : End R M} {μ : R} {v : M} (hv : f.HasEigenvector μ v) (n : ℕ) : - (f ^ n) v = μ ^ n • v := by - induction n <;> simp [*, pow_succ f, hv.apply_eq_smul, smul_smul, pow_succ' μ] + (f ^ n) v = μ ^ n • v := + hv.pow_apply n theorem HasEigenvalue.exists_hasEigenvector {f : End R M} {μ : R} (hμ : f.HasEigenvalue μ) : ∃ v, f.HasEigenvector μ v := Submodule.exists_mem_ne_zero_of_ne_bot hμ +nonrec lemma HasEigenvalue.pow {f : End R M} {μ : R} (h : f.HasEigenvalue μ) (n : ℕ) : - (f ^ n).HasEigenvalue (μ ^ n) := by - rw [HasEigenvalue, Submodule.ne_bot_iff] - obtain ⟨m : M, hm⟩ := h.exists_hasEigenvector - exact ⟨m, by simpa [mem_eigenspace_iff] using hm.pow_apply n, hm.2⟩ + (f ^ n).HasEigenvalue (μ ^ n) := + h.pow n /-- A nilpotent endomorphism has nilpotent eigenvalues. See also `LinearMap.isNilpotent_trace_of_isNilpotent`. -/ +nonrec lemma HasEigenvalue.isNilpotent_of_isNilpotent [NoZeroSMulDivisors R M] {f : End R M} (hfn : IsNilpotent f) {μ : R} (hf : f.HasEigenvalue μ) : - IsNilpotent μ := by - obtain ⟨m : M, hm⟩ := hf.exists_hasEigenvector - obtain ⟨n : ℕ, hn : f ^ n = 0⟩ := hfn - exact ⟨n, by simpa [hn, hm.2, eq_comm (a := (0 : M))] using hm.pow_apply n⟩ + IsNilpotent μ := + hf.isNilpotent_of_isNilpotent hfn +nonrec theorem HasEigenvalue.mem_spectrum {f : End R M} {μ : R} (hμ : HasEigenvalue f μ) : - μ ∈ spectrum R f := by - refine spectrum.mem_iff.mpr fun h_unit => ?_ - set f' := LinearMap.GeneralLinearGroup.toLinearEquiv h_unit.unit - rcases hμ.exists_hasEigenvector with ⟨v, hv⟩ - refine hv.2 ((LinearMap.ker_eq_bot'.mp f'.ker) v (?_ : μ • v - f v = 0)) - rw [hv.apply_eq_smul, sub_self] + μ ∈ spectrum R f := + hμ.mem_spectrum theorem hasEigenvalue_iff_mem_spectrum [FiniteDimensional K V] {f : End K V} {μ : K} : - f.HasEigenvalue μ ↔ μ ∈ spectrum K f := by - rw [spectrum.mem_iff, IsUnit.sub_iff, LinearMap.isUnit_iff_ker_eq_bot, HasEigenvalue, eigenspace] + f.HasEigenvalue μ ↔ μ ∈ spectrum K f := + hasUnifEigenvalue_iff_mem_spectrum alias ⟨_, HasEigenvalue.of_mem_spectrum⟩ := hasEigenvalue_iff_mem_spectrum theorem eigenspace_div (f : End K V) (a b : K) (hb : b ≠ 0) : eigenspace f (a / b) = LinearMap.ker (b • f - algebraMap K (End K V) a) := - calc - eigenspace f (a / b) = eigenspace f (b⁻¹ * a) := by rw [div_eq_mul_inv, mul_comm] - _ = LinearMap.ker (f - (b⁻¹ * a) • LinearMap.id) := by rw [eigenspace]; rfl - _ = LinearMap.ker (f - b⁻¹ • a • LinearMap.id) := by rw [smul_smul] - _ = LinearMap.ker (f - b⁻¹ • algebraMap K (End K V) a) := rfl - _ = LinearMap.ker (b • (f - b⁻¹ • algebraMap K (End K V) a)) := by - rw [LinearMap.ker_smul _ b hb] - _ = LinearMap.ker (b • f - algebraMap K (End K V) a) := by rw [smul_sub, smul_inv_smul₀ hb] + unifEigenspace_div f a b hb /-- The generalized eigenspace for a linear map `f`, a scalar `μ`, and an exponent `k ∈ ℕ` is the kernel of `(f - μ • id) ^ k`. (Def 8.10 of [axler2015]). Furthermore, a generalized eigenspace for some exponent `k` is contained in the generalized eigenspace for exponents larger than `k`. -/ def genEigenspace (f : End R M) (μ : R) : ℕ →o Submodule R M where - toFun k := LinearMap.ker ((f - algebraMap R (End R M) μ) ^ k) - monotone' k m hm := by - simp only [← pow_sub_mul_pow _ hm] - exact - LinearMap.ker_le_ker_comp ((f - algebraMap R (End R M) μ) ^ k) - ((f - algebraMap R (End R M) μ) ^ (m - k)) + toFun k := f.unifEigenspace μ k + monotone' k l hkl := (f.unifEigenspace μ).monotone <| by simpa lemma genEigenspace_def (f : End R M) (μ : R) (k : ℕ) : - f.genEigenspace μ k = LinearMap.ker ((f - algebraMap R (End R M) μ) ^ k) := rfl + f.genEigenspace μ k = LinearMap.ker ((f - μ • 1) ^ k) := by + rw [genEigenspace, OrderHom.coe_mk, unifEigenspace_nat] @[simp] theorem mem_genEigenspace (f : End R M) (μ : R) (k : ℕ) (m : M) : - m ∈ f.genEigenspace μ k ↔ ((f - μ • (1 : End R M)) ^ k) m = 0 := Iff.rfl + m ∈ f.genEigenspace μ k ↔ ((f - μ • (1 : End R M)) ^ k) m = 0 := + mem_unifEigenspace_nat @[simp] theorem genEigenspace_zero (f : End R M) (k : ℕ) : - f.genEigenspace 0 k = LinearMap.ker (f ^ k) := by - simp [Module.End.genEigenspace] + f.genEigenspace 0 k = LinearMap.ker (f ^ k) := + unifEigenspace_zero_nat _ _ /-- A nonzero element of a generalized eigenspace is a generalized eigenvector. (Def 8.9 of [axler2015])-/ -def HasGenEigenvector (f : End R M) (μ : R) (k : ℕ) (x : M) : Prop := - x ≠ 0 ∧ x ∈ genEigenspace f μ k +abbrev HasGenEigenvector (f : End R M) (μ : R) (k : ℕ) (x : M) : Prop := + HasUnifEigenvector f μ k x lemma hasGenEigenvector_iff {f : End R M} {μ : R} {k : ℕ} {x : M} : - f.HasGenEigenvector μ k x ↔ x ≠ 0 ∧ x ∈ f.genEigenspace μ k := Iff.rfl + f.HasGenEigenvector μ k x ↔ x ∈ f.genEigenspace μ k ∧ x ≠ 0 := Iff.rfl /-- A scalar `μ` is a generalized eigenvalue for a linear map `f` and an exponent `k ∈ ℕ` if there are generalized eigenvectors for `f`, `k`, and `μ`. -/ -def HasGenEigenvalue (f : End R M) (μ : R) (k : ℕ) : Prop := - genEigenspace f μ k ≠ ⊥ +abbrev HasGenEigenvalue (f : End R M) (μ : R) (k : ℕ) : Prop := + HasUnifEigenvalue f μ k -lemma hasGenEigenvalue_iff (f : End R M) (μ : R) (k : ℕ) : - f.HasGenEigenvalue μ k ↔ genEigenspace f μ k ≠ ⊥ := Iff.rfl +lemma hasGenEigenvalue_iff {f : End R M} {μ : R} {k : ℕ} : + f.HasGenEigenvalue μ k ↔ f.genEigenspace μ k ≠ ⊥ := Iff.rfl /-- The generalized eigenrange for a linear map `f`, a scalar `μ`, and an exponent `k ∈ ℕ` is the range of `(f - μ • id) ^ k`. -/ -def genEigenrange (f : End R M) (μ : R) (k : ℕ) : Submodule R M := - LinearMap.range ((f - algebraMap R (End R M) μ) ^ k) +abbrev genEigenrange (f : End R M) (μ : R) (k : ℕ) : Submodule R M := + unifEigenrange f μ k -lemma genEigenrange_def (f : End R M) (μ : R) (k : ℕ) : - f.genEigenrange μ k = LinearMap.range ((f - algebraMap R (End R M) μ) ^ k) := rfl +lemma genEigenrange_def {f : End R M} {μ : R} {k : ℕ} : + f.genEigenrange μ k = LinearMap.range ((f - μ • 1) ^ k) := by + rw [genEigenrange, unifEigenrange_nat] /-- The exponent of a generalized eigenvalue is never 0. -/ theorem exp_ne_zero_of_hasGenEigenvalue {f : End R M} {μ : R} {k : ℕ} - (h : f.HasGenEigenvalue μ k) : k ≠ 0 := by - rintro rfl - exact h LinearMap.ker_id + (h : f.HasGenEigenvalue μ k) : k ≠ 0 := + HasUnifEigenvalue.exp_ne_zero h /-- The union of the kernels of `(f - μ • id) ^ k` over all `k`. -/ -def maxGenEigenspace (f : End R M) (μ : R) : Submodule R M := - ⨆ k, f.genEigenspace μ k +abbrev maxGenEigenspace (f : End R M) (μ : R) : Submodule R M := + unifEigenspace f μ ⊤ lemma maxGenEigenspace_def (f : End R M) (μ : R) : - f.maxGenEigenspace μ = ⨆ k, f.genEigenspace μ k := rfl + f.maxGenEigenspace μ = ⨆ k, f.genEigenspace μ k := by + simp_rw [maxGenEigenspace, unifEigenspace_top, genEigenspace, OrderHom.coe_mk] theorem genEigenspace_le_maximal (f : End R M) (μ : R) (k : ℕ) : f.genEigenspace μ k ≤ f.maxGenEigenspace μ := - le_iSup _ _ + (f.unifEigenspace μ).monotone le_top @[simp] theorem mem_maxGenEigenspace (f : End R M) (μ : R) (m : M) : - m ∈ f.maxGenEigenspace μ ↔ ∃ k : ℕ, ((f - μ • (1 : End R M)) ^ k) m = 0 := by - simp only [maxGenEigenspace, ← mem_genEigenspace, Submodule.mem_iSup_of_chain] + m ∈ f.maxGenEigenspace μ ↔ ∃ k : ℕ, ((f - μ • (1 : End R M)) ^ k) m = 0 := + mem_unifEigenspace_top /-- If there exists a natural number `k` such that the kernel of `(f - μ • id) ^ k` is the maximal generalized eigenspace, then this value is the least such `k`. If not, this value is not meaningful. -/ -noncomputable def maxGenEigenspaceIndex (f : End R M) (μ : R) := - monotonicSequenceLimitIndex (f.genEigenspace μ) +noncomputable abbrev maxGenEigenspaceIndex (f : End R M) (μ : R) := + maxUnifEigenspaceIndex f μ /-- For an endomorphism of a Noetherian module, the maximal eigenspace is always of the form kernel `(f - μ • id) ^ k` for some `k`. -/ -theorem maxGenEigenspace_eq [h : IsNoetherian R M] (f : End R M) (μ : R) : +theorem maxGenEigenspace_eq [IsNoetherian R M] (f : End R M) (μ : R) : maxGenEigenspace f μ = f.genEigenspace μ (maxGenEigenspaceIndex f μ) := - h.wf.iSup_eq_monotonicSequenceLimit (f.genEigenspace μ) + unifEigenspace_top_eq_maxUnifEigenspaceIndex _ _ /-- A generalized eigenvalue for some exponent `k` is also a generalized eigenvalue for exponents larger than `k`. -/ theorem hasGenEigenvalue_of_hasGenEigenvalue_of_le {f : End R M} {μ : R} {k : ℕ} {m : ℕ} (hm : k ≤ m) (hk : f.HasGenEigenvalue μ k) : - f.HasGenEigenvalue μ m := by - unfold HasGenEigenvalue at * - contrapose! hk - rw [← le_bot_iff, ← hk] - exact (f.genEigenspace μ).monotone hm + f.HasGenEigenvalue μ m := + hk.le <| by simpa using hm /-- The eigenspace is a subspace of the generalized eigenspace. -/ theorem eigenspace_le_genEigenspace {f : End R M} {μ : R} {k : ℕ} (hk : 0 < k) : f.eigenspace μ ≤ f.genEigenspace μ k := - (f.genEigenspace μ).monotone (Nat.succ_le_of_lt hk) + (f.unifEigenspace _).monotone <| by simpa using Nat.succ_le_of_lt hk /-- All eigenvalues are generalized eigenvalues. -/ theorem hasGenEigenvalue_of_hasEigenvalue {f : End R M} {μ : R} {k : ℕ} (hk : 0 < k) - (hμ : f.HasEigenvalue μ) : f.HasGenEigenvalue μ k := by - apply hasGenEigenvalue_of_hasGenEigenvalue_of_le hk - rw [HasGenEigenvalue, genEigenspace, OrderHom.coe_mk, pow_one] - exact hμ + (hμ : f.HasEigenvalue μ) : f.HasGenEigenvalue μ k := + hμ.lt <| by simpa using hk /-- All generalized eigenvalues are eigenvalues. -/ theorem hasEigenvalue_of_hasGenEigenvalue {f : End R M} {μ : R} {k : ℕ} - (hμ : f.HasGenEigenvalue μ k) : f.HasEigenvalue μ := by - intro contra; apply hμ - erw [LinearMap.ker_eq_bot] at contra ⊢; rw [LinearMap.coe_pow] - exact Function.Injective.iterate contra k + (hμ : f.HasGenEigenvalue μ k) : f.HasEigenvalue μ := + hμ.lt zero_lt_one /-- Generalized eigenvalues are actually just eigenvalues. -/ @[simp] theorem hasGenEigenvalue_iff_hasEigenvalue {f : End R M} {μ : R} {k : ℕ} (hk : 0 < k) : f.HasGenEigenvalue μ k ↔ f.HasEigenvalue μ := - ⟨hasEigenvalue_of_hasGenEigenvalue, hasGenEigenvalue_of_hasEigenvalue hk⟩ + hasUnifEigenvalue_iff_hasUnifEigenvalue_one <| by simpa using hk /-- Every generalized eigenvector is a generalized eigenvector for exponent `finrank K V`. (Lemma 8.11 of [axler2015]) -/ theorem genEigenspace_le_genEigenspace_finrank [FiniteDimensional K V] (f : End K V) (μ : K) (k : ℕ) : f.genEigenspace μ k ≤ f.genEigenspace μ (finrank K V) := - ker_pow_le_ker_pow_finrank _ _ + unifEigenspace_le_unifEigenspace_finrank _ _ _ @[simp] theorem iSup_genEigenspace_eq_genEigenspace_finrank [FiniteDimensional K V] (f : End K V) (μ : K) : @@ -289,7 +587,7 @@ theorem genEigenspace_le_genEigenspace_finrank [FiniteDimensional K V] (f : End theorem genEigenspace_eq_genEigenspace_finrank_of_le [FiniteDimensional K V] (f : End K V) (μ : K) {k : ℕ} (hk : finrank K V ≤ k) : f.genEigenspace μ k = f.genEigenspace μ (finrank K V) := - ker_pow_eq_ker_pow_finrank_of_le hk + unifEigenspace_eq_unifEigenspace_finrank_of_le f μ hk lemma mapsTo_genEigenspace_of_comm {f g : End R M} (h : Commute f g) (μ : R) (k : ℕ) : MapsTo g (f.genEigenspace μ k) (f.genEigenspace μ k) := by @@ -300,21 +598,43 @@ lemma mapsTo_genEigenspace_of_comm {f g : End R M} (h : Commute f g) (μ : R) (k rw [← LinearMap.comp_apply, ← LinearMap.mul_eq_comp, h.eq, LinearMap.mul_eq_comp, LinearMap.comp_apply, hx, map_zero] -lemma mapsTo_iSup_genEigenspace_of_comm {f g : End R M} (h : Commute f g) (μ : R) : - MapsTo g ↑(⨆ k, f.genEigenspace μ k) ↑(⨆ k, f.genEigenspace μ k) := by +lemma iSup_genEigenspace_eq (f : End R M) (μ : R) : + ⨆ k, (f.genEigenspace μ) k = f.unifEigenspace μ ⊤ := by + rw [unifEigenspace_eq_iSup_unifEigenspace_nat] + ext + simp only [iSup_subtype, le_top, iSup_pos] + rfl + +lemma mapsTo_maxGenEigenspace_of_comm {f g : End R M} (h : Commute f g) (μ : R) : + MapsTo g ↑(f.maxGenEigenspace μ) ↑(f.maxGenEigenspace μ) := by + rw [maxGenEigenspace_def] simp only [MapsTo, Submodule.coe_iSup_of_chain, mem_iUnion, SetLike.mem_coe] rintro x ⟨k, hk⟩ exact ⟨k, f.mapsTo_genEigenspace_of_comm h μ k hk⟩ +lemma mapsTo_iSup_genEigenspace_of_comm {f g : End R M} (h : Commute f g) (μ : R) : + MapsTo g ↑(⨆ k, f.genEigenspace μ k) ↑(⨆ k, f.genEigenspace μ k) := by + rw [← maxGenEigenspace_def] + apply mapsTo_maxGenEigenspace_of_comm h + /-- The restriction of `f - μ • 1` to the `k`-fold generalized `μ`-eigenspace is nilpotent. -/ lemma isNilpotent_restrict_sub_algebraMap (f : End R M) (μ : R) (k : ℕ) (h : MapsTo (f - algebraMap R (End R M) μ) (f.genEigenspace μ k) (f.genEigenspace μ k) := mapsTo_genEigenspace_of_comm (Algebra.mul_sub_algebraMap_commutes f μ) μ k) : + IsNilpotent ((f - algebraMap R (End R M) μ).restrict h) := + isNilpotent_restrict_unifEigenspace_nat _ _ _ + +/-- The restriction of `f - μ • 1` to the generalized `μ`-eigenspace is nilpotent. -/ +lemma isNilpotent_restrict_maxGenEigenspace_sub_algebraMap [IsNoetherian R M] (f : End R M) (μ : R) + (h : MapsTo (f - algebraMap R (End R M) μ) + ↑(f.maxGenEigenspace μ) ↑(f.maxGenEigenspace μ) := + mapsTo_maxGenEigenspace_of_comm (Algebra.mul_sub_algebraMap_commutes f μ) μ) : IsNilpotent ((f - algebraMap R (End R M) μ).restrict h) := by - use k - ext - simp [LinearMap.restrict_apply, LinearMap.pow_restrict _] + apply isNilpotent_restrict_of_le (q := f.unifEigenspace μ (maxUnifEigenspaceIndex f μ)) + _ (isNilpotent_restrict_unifEigenspace_nat f μ (maxUnifEigenspaceIndex f μ)) + rw [maxGenEigenspace_eq] + exact le_rfl /-- The restriction of `f - μ • 1` to the generalized `μ`-eigenspace is nilpotent. -/ lemma isNilpotent_restrict_iSup_sub_algebraMap [IsNoetherian R M] (f : End R M) (μ : R) @@ -322,42 +642,53 @@ lemma isNilpotent_restrict_iSup_sub_algebraMap [IsNoetherian R M] (f : End R M) ↑(⨆ k, f.genEigenspace μ k) ↑(⨆ k, f.genEigenspace μ k) := mapsTo_iSup_genEigenspace_of_comm (Algebra.mul_sub_algebraMap_commutes f μ) μ) : IsNilpotent ((f - algebraMap R (End R M) μ).restrict h) := by - obtain ⟨l, hl⟩ : ∃ l, ⨆ k, f.genEigenspace μ k = f.genEigenspace μ l := - ⟨_, maxGenEigenspace_eq f μ⟩ - use l - ext ⟨x, hx⟩ - simpa [hl, LinearMap.restrict_apply, LinearMap.pow_restrict _] using hx - -lemma disjoint_genEigenspace [NoZeroSMulDivisors R M] - (f : End R M) {μ₁ μ₂ : R} (hμ : μ₁ ≠ μ₂) (k l : ℕ) : - Disjoint (f.genEigenspace μ₁ k) (f.genEigenspace μ₂ l) := by + apply isNilpotent_restrict_of_le (q := f.unifEigenspace μ (maxUnifEigenspaceIndex f μ)) + _ (isNilpotent_restrict_unifEigenspace_nat f μ (maxUnifEigenspaceIndex f μ)) + apply iSup_le + intro k + apply unifEigenspace_le_unifEigenspace_maxUnifEigenspaceIndex + +lemma disjoint_unifEigenspace [NoZeroSMulDivisors R M] + (f : End R M) {μ₁ μ₂ : R} (hμ : μ₁ ≠ μ₂) (k l : ℕ∞) : + Disjoint (f.unifEigenspace μ₁ k) (f.unifEigenspace μ₂ l) := by + rw [unifEigenspace_eq_iSup_unifEigenspace_nat, unifEigenspace_eq_iSup_unifEigenspace_nat] + simp_rw [unifEigenspace_directed.disjoint_iSup_left, unifEigenspace_directed.disjoint_iSup_right] + rintro ⟨k, -⟩ ⟨l, -⟩ nontriviality M have := NoZeroSMulDivisors.isReduced R M rw [disjoint_iff] - set p := f.genEigenspace μ₁ k ⊓ f.genEigenspace μ₂ l + set p := f.unifEigenspace μ₁ k ⊓ f.unifEigenspace μ₂ l by_contra hp replace hp : Nontrivial p := Submodule.nontrivial_iff_ne_bot.mpr hp let f₁ : End R p := (f - algebraMap R (End R M) μ₁).restrict <| MapsTo.inter_inter - (mapsTo_genEigenspace_of_comm (Algebra.mul_sub_algebraMap_commutes f μ₁) μ₁ k) - (mapsTo_genEigenspace_of_comm (Algebra.mul_sub_algebraMap_commutes f μ₁) μ₂ l) + (mapsTo_unifEigenspace_of_comm (Algebra.mul_sub_algebraMap_commutes f μ₁) μ₁ k) + (mapsTo_unifEigenspace_of_comm (Algebra.mul_sub_algebraMap_commutes f μ₁) μ₂ l) let f₂ : End R p := (f - algebraMap R (End R M) μ₂).restrict <| MapsTo.inter_inter - (mapsTo_genEigenspace_of_comm (Algebra.mul_sub_algebraMap_commutes f μ₂) μ₁ k) - (mapsTo_genEigenspace_of_comm (Algebra.mul_sub_algebraMap_commutes f μ₂) μ₂ l) + (mapsTo_unifEigenspace_of_comm (Algebra.mul_sub_algebraMap_commutes f μ₂) μ₁ k) + (mapsTo_unifEigenspace_of_comm (Algebra.mul_sub_algebraMap_commutes f μ₂) μ₂ l) have : IsNilpotent (f₂ - f₁) := by - apply Commute.isNilpotent_sub (x := f₂) (y := f₁) _ ⟨l, ?_⟩ ⟨k, ?_⟩ + apply Commute.isNilpotent_sub (x := f₂) (y := f₁) _ + (isNilpotent_restrict_of_le inf_le_right _) + (isNilpotent_restrict_of_le inf_le_left _) · ext; simp [f₁, f₂, smul_sub, sub_sub, smul_comm μ₁, add_sub_left_comm] - all_goals ext ⟨x, _, _⟩; simpa [LinearMap.restrict_apply, LinearMap.pow_restrict _] using ‹_› + apply mapsTo_unifEigenspace_of_comm (Algebra.mul_sub_algebraMap_commutes f _) + apply isNilpotent_restrict_unifEigenspace_nat + apply mapsTo_unifEigenspace_of_comm (Algebra.mul_sub_algebraMap_commutes f _) + apply isNilpotent_restrict_unifEigenspace_nat have hf₁₂ : f₂ - f₁ = algebraMap R (End R p) (μ₁ - μ₂) := by ext; simp [f₁, f₂, sub_smul] rw [hf₁₂, IsNilpotent.map_iff (NoZeroSMulDivisors.algebraMap_injective R (End R p)), isNilpotent_iff_eq_zero, sub_eq_zero] at this contradiction +lemma disjoint_genEigenspace [NoZeroSMulDivisors R M] + (f : End R M) {μ₁ μ₂ : R} (hμ : μ₁ ≠ μ₂) (k l : ℕ) : + Disjoint (f.genEigenspace μ₁ k) (f.genEigenspace μ₂ l) := + disjoint_unifEigenspace f hμ k l + lemma disjoint_iSup_genEigenspace [NoZeroSMulDivisors R M] (f : End R M) {μ₁ μ₂ : R} (hμ : μ₁ ≠ μ₂) : Disjoint (⨆ k, f.genEigenspace μ₁ k) (⨆ k, f.genEigenspace μ₂ k) := by - simp_rw [(f.genEigenspace μ₁).mono.directed_le.disjoint_iSup_left, - (f.genEigenspace μ₂).mono.directed_le.disjoint_iSup_right] - exact disjoint_genEigenspace f hμ + simpa only [iSup_genEigenspace_eq] using disjoint_unifEigenspace f hμ ⊤ ⊤ lemma injOn_genEigenspace [NoZeroSMulDivisors R M] (f : End R M) : InjOn (⨆ k, f.genEigenspace · k) {μ | ⨆ k, f.genEigenspace μ k ≠ ⊥} := by @@ -366,12 +697,14 @@ lemma injOn_genEigenspace [NoZeroSMulDivisors R M] (f : End R M) : apply hμ₂ simpa only [hμ₁₂, disjoint_self] using f.disjoint_iSup_genEigenspace contra -theorem independent_genEigenspace [NoZeroSMulDivisors R M] (f : End R M) : - CompleteLattice.Independent (fun μ ↦ ⨆ k, f.genEigenspace μ k) := by +theorem independent_maxGenEigenspace [NoZeroSMulDivisors R M] (f : End R M) : + CompleteLattice.Independent f.maxGenEigenspace := by classical suffices ∀ μ (s : Finset R), μ ∉ s → Disjoint (⨆ k, f.genEigenspace μ k) (s.sup fun μ ↦ ⨆ k, f.genEigenspace μ k) by - simp_rw [CompleteLattice.independent_iff_supIndep_of_injOn f.injOn_genEigenspace, + show CompleteLattice.Independent (f.maxGenEigenspace ·) + simp_rw [maxGenEigenspace_def, + CompleteLattice.independent_iff_supIndep_of_injOn f.injOn_genEigenspace, Finset.supIndep_iff_disjoint_erase] exact fun s μ _ ↦ this _ _ (s.not_mem_erase μ) intro μ₁ s @@ -405,6 +738,11 @@ theorem independent_genEigenspace [NoZeroSMulDivisors R M] (f : End R M) : simp_rw [Submodule.mem_iSup_of_chain, mem_genEigenspace] exact ⟨k, hyz⟩ +theorem independent_genEigenspace [NoZeroSMulDivisors R M] (f : End R M) : + CompleteLattice.Independent (fun μ ↦ ⨆ k, f.genEigenspace μ k) := by + simp_rw [← maxGenEigenspace_def] + apply independent_maxGenEigenspace + /-- The eigenspaces of a linear operator form an independent family of subspaces of `M`. That is, any eigenspace has trivial intersection with the span of all the other eigenspaces. -/ theorem eigenspaces_independent [NoZeroSMulDivisors R M] (f : End R M) : @@ -417,7 +755,7 @@ theorem eigenvectors_linearIndependent' {ι : Type*} [NoZeroSMulDivisors R M] (f : End R M) (μ : ι → R) (hμ : Function.Injective μ) (v : ι → M) (h_eigenvec : ∀ i, f.HasEigenvector (μ i) (v i)) : LinearIndependent R v := f.eigenspaces_independent.comp hμ |>.linearIndependent _ - (fun i => h_eigenvec i |>.left) (fun i => h_eigenvec i |>.right) + (fun i ↦ h_eigenvec i |>.left) (fun i ↦ h_eigenvec i |>.right) /-- Eigenvectors corresponding to distinct eigenvalues of a linear operator are linearly independent. (Lemma 5.10 of [axler2015]) @@ -428,7 +766,7 @@ theorem eigenvectors_linearIndependent' {ι : Type*} [NoZeroSMulDivisors R M] theorem eigenvectors_linearIndependent [NoZeroSMulDivisors R M] (f : End R M) (μs : Set R) (xs : μs → M) (h_eigenvec : ∀ μ : μs, f.HasEigenvector μ (xs μ)) : LinearIndependent R xs := - f.eigenvectors_linearIndependent' (fun μ : μs => μ) Subtype.coe_injective _ h_eigenvec + f.eigenvectors_linearIndependent' (fun μ : μs ↦ μ) Subtype.coe_injective _ h_eigenvec /-- If `f` maps a subspace `p` into itself, then the generalized eigenspace of the restriction of `f` to `p` is the part of the generalized eigenspace of `f` that lies in `p`. -/ @@ -436,7 +774,7 @@ theorem genEigenspace_restrict (f : End R M) (p : Submodule R M) (k : ℕ) (μ : (hfp : ∀ x : M, x ∈ p → f x ∈ p) : genEigenspace (LinearMap.restrict f hfp) μ k = Submodule.comap p.subtype (f.genEigenspace μ k) := by - simp only [genEigenspace, OrderHom.coe_mk, ← LinearMap.ker_comp] + simp only [genEigenspace_def, OrderHom.coe_mk, ← LinearMap.ker_comp] induction' k with k ih · rw [pow_zero, pow_zero, LinearMap.one_eq_id] apply (Submodule.ker_subtype _).symm @@ -459,7 +797,7 @@ lemma _root_.Submodule.inf_iInf_maxGenEigenspace_of_forall_mapsTo {ι : Type*} { (⨅ i, maxGenEigenspace ((f i).restrict (hfp i)) (μ i)).map p.subtype := by cases isEmpty_or_nonempty ι · simp [iInf_of_isEmpty] - · simp_rw [inf_iInf, maxGenEigenspace, ((f _).genEigenspace _).mono.directed_le.inf_iSup_eq, + · simp_rw [inf_iInf, maxGenEigenspace_def, ((f _).genEigenspace _).mono.directed_le.inf_iSup_eq, p.inf_genEigenspace _ (hfp _), ← Submodule.map_iSup, Submodule.map_iInf _ p.injective_subtype] /-- Given a family of endomorphisms `i ↦ f i`, a family of candidate eigenvalues `i ↦ μ i`, and a @@ -480,7 +818,6 @@ lemma iInf_maxGenEigenspace_restrict_map_subtype_eq rw [Submodule.map_iInf _ p.injective_subtype, this, Submodule.inf_iInf] simp_rw [maxGenEigenspace_def, Submodule.map_iSup, ((f _).genEigenspace _).mono.directed_le.inf_iSup_eq, p.inf_genEigenspace (f _) (h _)] - rfl lemma mapsTo_restrict_maxGenEigenspace_restrict_of_mapsTo {p : Submodule R M} (f g : End R M) (hf : MapsTo f p p) (hg : MapsTo g p p) {μ₁ μ₂ : R} @@ -508,16 +845,18 @@ theorem generalized_eigenvec_disjoint_range_ker [FiniteDimensional K V] (f : End (f.genEigenspace μ (finrank K V)) := by have h := calc - Submodule.comap ((f - algebraMap _ _ μ) ^ finrank K V) + Submodule.comap ((f - μ • 1) ^ finrank K V) (f.genEigenspace μ (finrank K V)) = LinearMap.ker ((f - algebraMap _ _ μ) ^ finrank K V * (f - algebraMap K (End K V) μ) ^ finrank K V) := by - rw [genEigenspace, OrderHom.coe_mk, ← LinearMap.ker_comp]; rfl - _ = f.genEigenspace μ (finrank K V + finrank K V) := by rw [← pow_add]; rfl + rw [genEigenspace, OrderHom.coe_mk, unifEigenspace_nat, ← LinearMap.ker_comp]; rfl + _ = f.genEigenspace μ (finrank K V + finrank K V) := by + rw [← pow_add, genEigenspace, OrderHom.coe_mk, unifEigenspace_nat]; rfl _ = f.genEigenspace μ (finrank K V) := by - rw [genEigenspace_eq_genEigenspace_finrank_of_le]; omega - rw [disjoint_iff_inf_le, genEigenrange, LinearMap.range_eq_map, - Submodule.map_inf_eq_map_inf_comap, top_inf_eq, h] + rw [genEigenspace_eq_genEigenspace_finrank_of_le]; omega + rw [disjoint_iff_inf_le, genEigenrange, unifEigenrange_nat, LinearMap.range_eq_map, + Submodule.map_inf_eq_map_inf_comap, top_inf_eq, h, + genEigenspace, OrderHom.coe_mk, unifEigenspace_nat] apply Submodule.map_comap_le /-- If an invariant subspace `p` of an endomorphism `f` is disjoint from the `μ`-eigenspace of `f`, @@ -544,11 +883,13 @@ theorem map_genEigenrange_le {f : End K V} {μ : K} {n : ℕ} : calc Submodule.map f (f.genEigenrange μ n) = LinearMap.range (f * (f - algebraMap _ _ μ) ^ n) := by - rw [genEigenrange]; exact (LinearMap.range_comp _ _).symm + rw [genEigenrange, unifEigenrange_nat]; exact (LinearMap.range_comp _ _).symm _ = LinearMap.range ((f - algebraMap _ _ μ) ^ n * f) := by rw [Algebra.mul_sub_algebraMap_pow_commutes] _ = Submodule.map ((f - algebraMap _ _ μ) ^ n) (LinearMap.range f) := LinearMap.range_comp _ _ - _ ≤ f.genEigenrange μ n := LinearMap.map_le_range + _ ≤ f.genEigenrange μ n := by + rw [genEigenrange, unifEigenrange_nat] + apply LinearMap.map_le_range lemma iSup_genEigenspace_le_smul (f : Module.End R M) (μ t : R) : (⨆ k, f.genEigenspace μ k) ≤ ⨆ k, (t • f).genEigenspace (t * μ) k := by diff --git a/Mathlib/LinearAlgebra/Eigenspace/Minpoly.lean b/Mathlib/LinearAlgebra/Eigenspace/Minpoly.lean index e6cf66f99d9ea..70c17ee7b2097 100644 --- a/Mathlib/LinearAlgebra/Eigenspace/Minpoly.lean +++ b/Mathlib/LinearAlgebra/Eigenspace/Minpoly.lean @@ -78,7 +78,7 @@ theorem hasEigenvalue_of_isRoot (h : (minpoly K f).IsRoot μ) : f.HasEigenvalue have : (aeval f) p = 0 := by have h_aeval := minpoly.aeval K f revert h_aeval - simp [hp, ← hu] + simp [hp, ← hu, Algebra.algebraMap_eq_smul_one] have h_deg := minpoly.degree_le_of_ne_zero K f p_ne_0 this rw [hp, degree_mul, degree_X_sub_C, Polynomial.degree_eq_natDegree p_ne_0] at h_deg norm_cast at h_deg diff --git a/Mathlib/LinearAlgebra/Eigenspace/Triangularizable.lean b/Mathlib/LinearAlgebra/Eigenspace/Triangularizable.lean index c0f53eccf53f8..03e4ac0132b66 100644 --- a/Mathlib/LinearAlgebra/Eigenspace/Triangularizable.lean +++ b/Mathlib/LinearAlgebra/Eigenspace/Triangularizable.lean @@ -250,7 +250,8 @@ lemma Module.End.iSup_iInf_maxGenEigenspace_eq_top_of_forall_mapsTo apply ih _ (hy φ) · intro j k μ exact mapsTo_restrict_maxGenEigenspace_restrict_of_mapsTo (f j) (f k) _ _ (h j k μ) - · exact fun j ↦ Module.End.iSup_genEigenspace_restrict_eq_top _ (h' j) + · simp_rw [maxGenEigenspace_def] at h' ⊢ + exact fun j ↦ Module.End.iSup_genEigenspace_restrict_eq_top _ (h' j) · rfl replace ih (φ : K) : ⨆ (χ : ι → K) (_ : χ i = φ), ⨅ j, maxGenEigenspace ((f j).restrict (hi j φ)) (χ j) = ⊤ := by @@ -260,6 +261,7 @@ lemma Module.End.iSup_iInf_maxGenEigenspace_eq_top_of_forall_mapsTo rw [eq_bot_iff, ← ((f i).maxGenEigenspace φ).ker_subtype, LinearMap.ker, ← Submodule.map_le_iff_le_comap, ← Submodule.inf_iInf_maxGenEigenspace_of_forall_mapsTo, ← disjoint_iff_inf_le] + simp_rw [maxGenEigenspace_def] exact ((f i).disjoint_iSup_genEigenspace hχ.symm).mono_right (iInf_le _ i) replace ih (φ : K) : ⨆ (χ : ι → K) (_ : χ i = φ), ⨅ j, maxGenEigenspace (f j) (χ j) = diff --git a/Mathlib/Order/Hom/Basic.lean b/Mathlib/Order/Hom/Basic.lean index f77c690521b89..356b8c434a945 100644 --- a/Mathlib/Order/Hom/Basic.lean +++ b/Mathlib/Order/Hom/Basic.lean @@ -1172,6 +1172,13 @@ theorem coe_toDualTopEquiv_eq [LE α] : (WithBot.toDualTopEquiv : WithBot αᵒᵈ → (WithTop α)ᵒᵈ) = toDual ∘ WithBot.ofDual := funext fun _ => rfl +/-- The coercion `α → WithBot α` bundled as monotone map. -/ +@[simps] +def coeOrderHom {α : Type*} [Preorder α] : α ↪o WithBot α where + toFun := (↑) + inj' := WithBot.coe_injective + map_rel_iff' := WithBot.coe_le_coe + end WithBot namespace WithTop @@ -1203,6 +1210,13 @@ theorem coe_toDualBotEquiv [LE α] : (WithTop.toDualBotEquiv : WithTop αᵒᵈ → (WithBot α)ᵒᵈ) = toDual ∘ WithTop.ofDual := funext fun _ => rfl +/-- The coercion `α → WithTop α` bundled as monotone map. -/ +@[simps] +def coeOrderHom {α : Type*} [Preorder α] : α ↪o WithTop α where + toFun := (↑) + inj' := WithTop.coe_injective + map_rel_iff' := WithTop.coe_le_coe + end WithTop namespace OrderIso diff --git a/Mathlib/RingTheory/Nilpotent/Lemmas.lean b/Mathlib/RingTheory/Nilpotent/Lemmas.lean index a26835374afce..fdbae3b577172 100644 --- a/Mathlib/RingTheory/Nilpotent/Lemmas.lean +++ b/Mathlib/RingTheory/Nilpotent/Lemmas.lean @@ -89,13 +89,31 @@ end LinearMap namespace Module.End -lemma isNilpotent.restrict {R M : Type*} [Semiring R] [AddCommMonoid M] [Module R M] +section + +variable {M : Type*} [Semiring R] [AddCommMonoid M] [Module R M] + +lemma isNilpotent_restrict_of_le {f : End R M} {p q : Submodule R M} + {hp : MapsTo f p p} {hq : MapsTo f q q} (h : p ≤ q) (hf : IsNilpotent (f.restrict hq)) : + IsNilpotent (f.restrict hp) := by + obtain ⟨n, hn⟩ := hf + use n + ext ⟨x, hx⟩ + replace hn := DFunLike.congr_fun hn ⟨x, h hx⟩ + 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 : _) + +lemma isNilpotent.restrict {f : M →ₗ[R] M} {p : Submodule R M} (hf : MapsTo f p p) (hnil : IsNilpotent f) : IsNilpotent (f.restrict hf) := by obtain ⟨n, hn⟩ := hnil exact ⟨n, LinearMap.ext fun m ↦ by simp only [LinearMap.pow_restrict n, hn, LinearMap.restrict_apply, LinearMap.zero_apply]; rfl⟩ +end + variable {M : Type v} [Ring R] [AddCommGroup M] [Module R M] variable {f : Module.End R M} {p : Submodule R M} (hp : p ≤ p.comap f) From af0da371151e731946f06c188c581a0d999b2bc1 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Mon, 7 Oct 2024 16:09:51 +0000 Subject: [PATCH 22/43] perf(NumberTheory): speed up some slow declarations (#17395) --- .../Cyclotomic/PrimitiveRoots.lean | 41 ++++++++----------- .../NumberField/Discriminant.lean | 15 ++++++- 2 files changed, 31 insertions(+), 25 deletions(-) diff --git a/Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean b/Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean index a9c07d9408500..e4a82b6a04edb 100644 --- a/Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean +++ b/Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean @@ -113,9 +113,10 @@ variable {C} /-- The `PowerBasis` given by a primitive root `η`. -/ @[simps!] protected noncomputable def powerBasis : PowerBasis K L := - PowerBasis.map (Algebra.adjoin.powerBasis <| (integral {n} K L).isIntegral ζ) <| - (Subalgebra.equivOfEq _ _ (IsCyclotomicExtension.adjoin_primitive_root_eq_top hζ)).trans - Subalgebra.topEquiv + -- this is purely an optimization + letI pb := Algebra.adjoin.powerBasis <| (integral {n} K L).isIntegral ζ + pb.map <| (Subalgebra.equivOfEq _ _ (IsCyclotomicExtension.adjoin_primitive_root_eq_top hζ)).trans + Subalgebra.topEquiv theorem powerBasis_gen_mem_adjoin_zeta_sub_one : (hζ.powerBasis K).gen ∈ adjoin K ({ζ - 1} : Set L) := by @@ -392,9 +393,8 @@ theorem norm_pow_sub_one_of_prime_pow_ne_two {k s : ℕ} (hζ : IsPrimitiveRoot [hpri : Fact (p : ℕ).Prime] [IsCyclotomicExtension {p ^ (k + 1)} K L] (hirr : Irreducible (cyclotomic (↑(p ^ (k + 1)) : ℕ) K)) (hs : s ≤ k) (htwo : p ^ (k - s + 1) ≠ 2) : norm K (ζ ^ (p : ℕ) ^ s - 1) = (p : K) ^ (p : ℕ) ^ s := by --- Porting note: `by simp` was `by linarith` that now fails. have hirr₁ : Irreducible (cyclotomic ((p : ℕ) ^ (k - s + 1)) K) := - cyclotomic_irreducible_pow_of_irreducible_pow hpri.1 (by simp) hirr + cyclotomic_irreducible_pow_of_irreducible_pow hpri.1 (by omega) hirr rw [← PNat.pow_coe] at hirr₁ set η := ζ ^ (p : ℕ) ^ s - 1 let η₁ : K⟮η⟯ := IntermediateField.AdjoinSimple.gen K η @@ -403,22 +403,17 @@ theorem norm_pow_sub_one_of_prime_pow_ne_two {k s : ℕ} (hζ : IsPrimitiveRoot refine IsPrimitiveRoot.pow (p ^ (k + 1)).pos hζ ?_ rw [PNat.pow_coe, ← pow_add, add_comm s, Nat.sub_add_cancel (le_trans hs (Nat.le_succ k))] have : IsCyclotomicExtension {p ^ (k - s + 1)} K K⟮η⟯ := by - suffices IsCyclotomicExtension {p ^ (k - s + 1)} K K⟮η + 1⟯.toSubalgebra by - have H : K⟮η + 1⟯.toSubalgebra = K⟮η⟯.toSubalgebra := by - simp only [IntermediateField.adjoin_simple_toSubalgebra_of_integral - ((integral {p ^ (k + 1)} K L).isIntegral _)] - refine Subalgebra.ext fun x => ⟨fun hx => adjoin_le ?_ hx, fun hx => adjoin_le ?_ hx⟩ - · simp only [Set.singleton_subset_iff, SetLike.mem_coe] - exact Subalgebra.add_mem _ (subset_adjoin (mem_singleton η)) (Subalgebra.one_mem _) - · simp only [Set.singleton_subset_iff, SetLike.mem_coe] - nth_rw 2 [← add_sub_cancel_right η 1] - exact Subalgebra.sub_mem _ (subset_adjoin (mem_singleton _)) (Subalgebra.one_mem _) --- Porting note: the previous proof was `rw [H] at this; exact this` but it now fails. - exact IsCyclotomicExtension.equiv _ _ _ (Subalgebra.equivOfEq _ _ H) --- Porting note: the next `refine` was `rw [H]`, abusing defeq, and it now fails. + have HKη : K⟮η⟯ = K⟮η + 1⟯ := by + refine le_antisymm ?_ ?_ + all_goals rw [IntermediateField.adjoin_simple_le_iff] + · nth_rw 2 [← add_sub_cancel_right η 1] + exact sub_mem (IntermediateField.mem_adjoin_simple_self K (η + 1)) (one_mem _) + · exact add_mem (IntermediateField.mem_adjoin_simple_self K η) (one_mem _) + rw [HKη] have H := IntermediateField.adjoin_simple_toSubalgebra_of_integral - ((integral {p ^ (k + 1)} K L).isIntegral (η + 1)) - refine @IsCyclotomicExtension.equiv _ _ _ _ _ _ _ _ _ ?_ (Subalgebra.equivOfEq _ _ H).symm + ((integral {p ^ (k + 1)} K L).isIntegral (η + 1)) + refine IsCyclotomicExtension.equiv _ _ _ (h := ?_) (.refl : K⟮η + 1⟯.toSubalgebra ≃ₐ[K] _) + rw [H] have hη' : IsPrimitiveRoot (η + 1) ↑(p ^ (k + 1 - s)) := by simpa using hη -- Porting note: `using 1` was not needed. convert hη'.adjoin_isCyclotomicExtension K using 1 @@ -427,10 +422,10 @@ theorem norm_pow_sub_one_of_prime_pow_ne_two {k s : ℕ} (hζ : IsPrimitiveRoot apply coe_submonoidClass_iff.1 convert hη using 1 rw [Nat.sub_add_comm hs, pow_coe] --- Porting note: the following `haveI` were not needed because the locale `cyclotomic` set them +-- Porting note: the following `have` were not needed because the locale `cyclotomic` set them -- as instances. - haveI := IsCyclotomicExtension.finiteDimensional {p ^ (k + 1)} K L - haveI := IsCyclotomicExtension.isGalois (p ^ (k + 1)) K L + have := IsCyclotomicExtension.finiteDimensional {p ^ (k + 1)} K L + have := IsCyclotomicExtension.isGalois (p ^ (k + 1)) K L rw [norm_eq_norm_adjoin K] have H := hη.sub_one_norm_isPrimePow ?_ hirr₁ htwo swap; · rw [PNat.pow_coe]; exact hpri.1.isPrimePow.pow (Nat.succ_ne_zero _) diff --git a/Mathlib/NumberTheory/NumberField/Discriminant.lean b/Mathlib/NumberTheory/NumberField/Discriminant.lean index b56a1d4722194..52a08d5551c4a 100644 --- a/Mathlib/NumberTheory/NumberField/Discriminant.lean +++ b/Mathlib/NumberTheory/NumberField/Discriminant.lean @@ -337,6 +337,8 @@ theorem finite_of_discr_bdd_of_isReal : (Set.finite_Icc (-C : ℤ) C)) (fun ⟨K, hK₀⟩ ⟨hK₁, hK₂⟩ ↦ ?_) -- We now need to prove that each field is generated by an element of the union of the rootset simp_rw [Set.mem_iUnion] + -- this is purely an optimization + have : CharZero K := SubsemiringClass.instCharZero K haveI : NumberField K := @NumberField.mk _ _ inferInstance hK₀ obtain ⟨w₀, hw₀⟩ := hK₁ suffices minkowskiBound K ↑1 < (convexBodyLTFactor K) * B by @@ -360,12 +362,15 @@ theorem finite_of_discr_bdd_of_isReal : · refine mem_rootSet.mpr ⟨minpoly.ne_zero hx, ?_⟩ exact (aeval_algebraMap_eq_zero_iff _ _ _).mpr (minpoly.aeval ℤ (x : K)) · rw [← (IntermediateField.lift_injective _).eq_iff, eq_comm] at hx₁ - convert hx₁ <;> simp + convert hx₁ + · simp only [IntermediateField.lift_top] + · simp only [IntermediateField.lift_adjoin, Set.image_singleton] have := one_le_convexBodyLTFactor K convert lt_of_le_of_lt (mul_right_mono (coe_le_coe.mpr this)) (ENNReal.mul_lt_mul_left' (by positivity) coe_ne_top (minkowskiBound_lt_boundOfDiscBdd hK₂)) simp_rw [ENNReal.coe_one, one_mul] + theorem finite_of_discr_bdd_of_isComplex : {K : { F : IntermediateField ℚ A // FiniteDimensional ℚ F} | haveI : NumberField K := @NumberField.mk _ _ inferInstance K.prop @@ -380,6 +385,8 @@ theorem finite_of_discr_bdd_of_isComplex : (Set.finite_Icc (-C : ℤ) C)) (fun ⟨K, hK₀⟩ ⟨hK₁, hK₂⟩ ↦ ?_) -- We now need to prove that each field is generated by an element of the union of the rootset simp_rw [Set.mem_iUnion] + -- this is purely an optimization + have : CharZero K := SubsemiringClass.instCharZero K haveI : NumberField K := @NumberField.mk _ _ inferInstance hK₀ obtain ⟨w₀, hw₀⟩ := hK₁ suffices minkowskiBound K ↑1 < (convexBodyLT'Factor K) * boundOfDiscBdd N by @@ -404,7 +411,9 @@ theorem finite_of_discr_bdd_of_isComplex : · refine mem_rootSet.mpr ⟨minpoly.ne_zero hx, ?_⟩ exact (aeval_algebraMap_eq_zero_iff _ _ _).mpr (minpoly.aeval ℤ (x : K)) · rw [← (IntermediateField.lift_injective _).eq_iff, eq_comm] at hx₁ - convert hx₁ <;> simp + convert hx₁ + · simp only [IntermediateField.lift_top] + · simp only [IntermediateField.lift_adjoin, Set.image_singleton] have := one_le_convexBodyLT'Factor K convert lt_of_le_of_lt (mul_right_mono (coe_le_coe.mpr this)) (ENNReal.mul_lt_mul_left' (by positivity) coe_ne_top (minkowskiBound_lt_boundOfDiscBdd hK₂)) @@ -419,6 +428,8 @@ theorem _root_.NumberField.finite_of_discr_bdd : refine Set.Finite.subset (Set.Finite.union (finite_of_discr_bdd_of_isReal A N) (finite_of_discr_bdd_of_isComplex A N)) ?_ rintro ⟨K, hK₀⟩ hK₁ + -- this is purely an optimization + have : CharZero K := SubsemiringClass.instCharZero K haveI : NumberField K := @NumberField.mk _ _ inferInstance hK₀ obtain ⟨w₀⟩ := (inferInstance : Nonempty (InfinitePlace K)) by_cases hw₀ : IsReal w₀ From db788ef25e0189285f1aadd85243f404d3771cc8 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 7 Oct 2024 16:09:52 +0000 Subject: [PATCH 23/43] =?UTF-8?q?feat:=20`lineMap=20p=E2=82=80=20(lineMap?= =?UTF-8?q?=20p=E2=82=80=20p=E2=82=81=20c)=20d=20=3D=20lineMap=20p?= =?UTF-8?q?=E2=82=80=20p=E2=82=81=20(d=20*=20c)`=20(#17462)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit From LeanCamCombi --- Mathlib/LinearAlgebra/AffineSpace/AffineMap.lean | 7 +++++++ 1 file changed, 7 insertions(+) diff --git a/Mathlib/LinearAlgebra/AffineSpace/AffineMap.lean b/Mathlib/LinearAlgebra/AffineSpace/AffineMap.lean index b66ea39b78a5a..eb8a41bbe9312 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/AffineMap.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/AffineMap.lean @@ -565,6 +565,13 @@ theorem lineMap_vsub_lineMap (p₁ p₂ p₃ p₄ : P1) (c : k) : lineMap p₁ p₂ c -ᵥ lineMap p₃ p₄ c = lineMap (p₁ -ᵥ p₃) (p₂ -ᵥ p₄) c := ((fst : P1 × P1 →ᵃ[k] P1) -ᵥ (snd : P1 × P1 →ᵃ[k] P1)).apply_lineMap (_, _) (_, _) c +@[simp] lemma lineMap_lineMap_right (p₀ p₁ : P1) (c d : k) : + lineMap p₀ (lineMap p₀ p₁ c) d = lineMap p₀ p₁ (d * c) := by simp [lineMap_apply, mul_smul] + +@[simp] lemma lineMap_lineMap_left (p₀ p₁ : P1) (c d : k) : + lineMap (lineMap p₀ p₁ c) p₁ d = lineMap p₀ p₁ (1 - (1 - d) * (1 - c)) := by + simp_rw [lineMap_apply_one_sub, ← lineMap_apply_one_sub p₁, lineMap_lineMap_right] + /-- Decomposition of an affine map in the special case when the point space and vector space are the same. -/ theorem decomp (f : V1 →ᵃ[k] V2) : (f : V1 → V2) = ⇑f.linear + fun _ => f 0 := by From 1b67a7c40b12c5d1bcf3e2c3a6a1b1e2767a87c5 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 7 Oct 2024 16:09:54 +0000 Subject: [PATCH 24/43] feat: `Fintype.balance` and complex numbers (#17464) From LeanAPAP --- Mathlib/Analysis/RCLike/Basic.lean | 9 +++++++++ Mathlib/Data/Complex/BigOperators.lean | 24 +++++++++++++++++++++++- 2 files changed, 32 insertions(+), 1 deletion(-) diff --git a/Mathlib/Analysis/RCLike/Basic.lean b/Mathlib/Analysis/RCLike/Basic.lean index 674e2c7b5df43..43c7971fb36a0 100644 --- a/Mathlib/Analysis/RCLike/Basic.lean +++ b/Mathlib/Analysis/RCLike/Basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Frédéric Dupuis -/ import Mathlib.Algebra.Algebra.Field +import Mathlib.Algebra.BigOperators.Balance import Mathlib.Algebra.Order.BigOperators.Expect import Mathlib.Algebra.Order.Star.Basic import Mathlib.Analysis.CStarAlgebra.Basic @@ -40,6 +41,7 @@ their counterparts in `Mathlib/Analysis/Complex/Basic.lean` (which causes linter A few lemmas requiring heavier imports are in `Mathlib/Data/RCLike/Lemmas.lean`. -/ +open Fintype open scoped BigOperators ComplexConjugate section @@ -239,6 +241,13 @@ instance (priority := 100) charZero_rclike : CharZero K := lemma ofReal_expect {α : Type*} (s : Finset α) (f : α → ℝ) : 𝔼 i ∈ s, f i = 𝔼 i ∈ s, (f i : K) := map_expect (algebraMap ..) .. +@[norm_cast] +lemma ofReal_balance {ι : Type*} [Fintype ι] (f : ι → ℝ) (i : ι) : + ((balance f i : ℝ) : K) = balance ((↑) ∘ f) i := map_balance (algebraMap ..) .. + +@[simp] lemma ofReal_comp_balance {ι : Type*} [Fintype ι] (f : ι → ℝ) : + ofReal ∘ balance f = balance (ofReal ∘ f : ι → K) := funext <| ofReal_balance _ + /-! ### The imaginary unit, `I` -/ /-- The imaginary unit. -/ diff --git a/Mathlib/Data/Complex/BigOperators.lean b/Mathlib/Data/Complex/BigOperators.lean index cd9cb5f344fa0..df65eec1fe310 100644 --- a/Mathlib/Data/Complex/BigOperators.lean +++ b/Mathlib/Data/Complex/BigOperators.lean @@ -3,13 +3,14 @@ Copyright (c) 2017 Kevin Buzzard. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kevin Buzzard, Mario Carneiro -/ -import Mathlib.Algebra.BigOperators.Expect +import Mathlib.Algebra.BigOperators.Balance import Mathlib.Data.Complex.Basic /-! # Finite sums and products of complex numbers -/ +open Fintype open scoped BigOperators namespace Complex @@ -28,6 +29,13 @@ theorem ofReal_sum (f : α → ℝ) : ((∑ i ∈ s, f i : ℝ) : ℂ) = ∑ i lemma ofReal_expect (f : α → ℝ) : (𝔼 i ∈ s, f i : ℝ) = 𝔼 i ∈ s, (f i : ℂ) := map_expect ofReal .. +@[simp, norm_cast] +lemma ofReal_balance [Fintype α] (f : α → ℝ) (a : α) : + ((balance f a : ℝ) : ℂ) = balance ((↑) ∘ f) a := by simp [balance] + +@[simp] lemma ofReal_comp_balance {ι : Type*} [Fintype ι] (f : ι → ℝ) : + ofReal ∘ balance f = balance (ofReal ∘ f : ι → ℂ) := funext <| ofReal_balance _ + @[simp] theorem re_sum (f : α → ℂ) : (∑ i ∈ s, f i).re = ∑ i ∈ s, (f i).re := map_sum reAddGroupHom f s @@ -36,6 +44,13 @@ theorem re_sum (f : α → ℂ) : (∑ i ∈ s, f i).re = ∑ i ∈ s, (f i).re lemma re_expect (f : α → ℂ) : (𝔼 i ∈ s, f i).re = 𝔼 i ∈ s, (f i).re := map_expect (LinearMap.mk reAddGroupHom.toAddHom (by simp)) f s +@[simp] +lemma re_balance [Fintype α] (f : α → ℂ) (a : α) : re (balance f a) = balance (re ∘ f) a := by + simp [balance] + +@[simp] lemma re_comp_balance {ι : Type*} [Fintype ι] (f : ι → ℂ) : + re ∘ balance f = balance (re ∘ f) := funext <| re_balance _ + @[simp] theorem im_sum (f : α → ℂ) : (∑ i ∈ s, f i).im = ∑ i ∈ s, (f i).im := map_sum imAddGroupHom f s @@ -44,4 +59,11 @@ theorem im_sum (f : α → ℂ) : (∑ i ∈ s, f i).im = ∑ i ∈ s, (f i).im lemma im_expect (f : α → ℂ) : (𝔼 i ∈ s, f i).im = 𝔼 i ∈ s, (f i).im := map_expect (LinearMap.mk imAddGroupHom.toAddHom (by simp)) f s +@[simp] +lemma im_balance [Fintype α] (f : α → ℂ) (a : α) : im (balance f a) = balance (im ∘ f) a := by + simp [balance] + +@[simp] lemma im_comp_balance {ι : Type*} [Fintype ι] (f : ι → ℂ) : + im ∘ balance f = balance (im ∘ f) := funext <| im_balance _ + end Complex From 8b37774dc16ac887bac897ed343fee37fb820216 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Mon, 7 Oct 2024 16:09:55 +0000 Subject: [PATCH 25/43] chore: process and remove adaptation_note for nightly-2024-09-06 (#17498) --- Mathlib/Algebra/Order/Floor.lean | 18 +++---------- Mathlib/Algebra/Ring/Parity.lean | 6 +---- .../SetFamily/KruskalKatona.lean | 12 ++------- .../SimpleGraph/Regularity/Equitabilise.lean | 6 +---- .../SimpleGraph/Triangle/Basic.lean | 6 +---- Mathlib/Data/Finset/NatDivisors.lean | 6 +---- Mathlib/Data/Int/Defs.lean | 7 +----- Mathlib/Data/List/Basic.lean | 7 +----- Mathlib/Data/List/Chain.lean | 25 +++---------------- Mathlib/Data/List/Nodup.lean | 18 +++---------- Mathlib/Data/List/Perm.lean | 6 +---- Mathlib/Data/List/Permutation.lean | 6 +---- Mathlib/Data/List/Rotate.lean | 6 +---- Mathlib/Data/Multiset/Basic.lean | 18 +++---------- Mathlib/Data/Nat/Factorization/Basic.lean | 6 +---- Mathlib/Data/Nat/Log.lean | 6 +---- Mathlib/Data/Nat/WithBot.lean | 6 +---- Mathlib/Data/PFunctor/Univariate/M.lean | 4 --- Mathlib/Logic/Encodable/Basic.lean | 6 +---- Mathlib/NumberTheory/Bertrand.lean | 6 +---- Mathlib/NumberTheory/Divisors.lean | 24 +++--------------- .../LegendreSymbol/GaussEisensteinLemmas.lean | 6 +---- Mathlib/NumberTheory/SumFourSquares.lean | 9 ++----- Mathlib/Order/Interval/Finset/Nat.lean | 8 ++---- .../Symmetric/NewtonIdentities.lean | 6 +---- .../Polynomial/Cyclotomic/Eval.lean | 9 ++----- 26 files changed, 41 insertions(+), 202 deletions(-) diff --git a/Mathlib/Algebra/Order/Floor.lean b/Mathlib/Algebra/Order/Floor.lean index 6bd51a8b46a96..478a21d5d6e8c 100644 --- a/Mathlib/Algebra/Order/Floor.lean +++ b/Mathlib/Algebra/Order/Floor.lean @@ -222,21 +222,13 @@ theorem floor_eq_zero : ⌊a⌋₊ = 0 ↔ a < 1 := by rw [← lt_one_iff, ← @cast_one α] exact floor_lt' Nat.one_ne_zero -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ theorem floor_eq_iff (ha : 0 ≤ a) : ⌊a⌋₊ = n ↔ ↑n ≤ a ∧ a < ↑n + 1 := by rw [← le_floor_iff ha, ← Nat.cast_one, ← Nat.cast_add, ← floor_lt ha, Nat.lt_add_one_iff, - le_antisymm_iff, _root_.and_comm] + le_antisymm_iff, and_comm] -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ theorem floor_eq_iff' (hn : n ≠ 0) : ⌊a⌋₊ = n ↔ ↑n ≤ a ∧ a < ↑n + 1 := by rw [← le_floor_iff' hn, ← Nat.cast_one, ← Nat.cast_add, ← floor_lt' (Nat.add_one_ne_zero n), - Nat.lt_add_one_iff, le_antisymm_iff, _root_.and_comm] + Nat.lt_add_one_iff, le_antisymm_iff, and_comm] theorem floor_eq_on_Ico (n : ℕ) : ∀ a ∈ (Set.Ico n (n + 1) : Set α), ⌊a⌋₊ = n := fun _ ⟨h₀, h₁⟩ => (floor_eq_iff <| n.cast_nonneg.trans h₀).mpr ⟨h₀, h₁⟩ @@ -328,14 +320,10 @@ theorem floor_lt_ceil_of_lt_of_pos {a b : α} (h : a < b) (h' : 0 < b) : ⌊a⌋ exact h.trans_le (le_ceil _) · rwa [floor_of_nonpos ha.le, lt_ceil, Nat.cast_zero] -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ theorem ceil_eq_iff (hn : n ≠ 0) : ⌈a⌉₊ = n ↔ ↑(n - 1) < a ∧ a ≤ n := by rw [← ceil_le, ← not_le, ← ceil_le, not_le, tsub_lt_iff_right (Nat.add_one_le_iff.2 (pos_iff_ne_zero.2 hn)), Nat.lt_add_one_iff, - le_antisymm_iff, _root_.and_comm] + le_antisymm_iff, and_comm] @[simp] theorem preimage_ceil_zero : (Nat.ceil : α → ℕ) ⁻¹' {0} = Iic 0 := diff --git a/Mathlib/Algebra/Ring/Parity.lean b/Mathlib/Algebra/Ring/Parity.lean index d257486efbda8..9205b9aeff65a 100644 --- a/Mathlib/Algebra/Ring/Parity.lean +++ b/Mathlib/Algebra/Ring/Parity.lean @@ -233,14 +233,10 @@ lemma even_or_odd (n : ℕ) : Even n ∨ Odd n := (even_xor_odd n).or lemma even_or_odd' (n : ℕ) : ∃ k, n = 2 * k ∨ n = 2 * k + 1 := by simpa only [← two_mul, exists_or, Odd, Even] using even_or_odd n -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ lemma even_xor_odd' (n : ℕ) : ∃ k, Xor' (n = 2 * k) (n = 2 * k + 1) := by obtain ⟨k, rfl⟩ | ⟨k, rfl⟩ := even_or_odd n <;> use k · simpa only [← two_mul, eq_self_iff_true, xor_true] using (succ_ne_self (2 * k)).symm - · simpa only [xor_true, _root_.xor_comm] using (succ_ne_self _) + · simpa only [xor_true, xor_comm] using (succ_ne_self _) lemma mod_two_add_add_odd_mod_two (m : ℕ) {n : ℕ} (hn : Odd n) : m % 2 + (m + n) % 2 = 1 := ((even_or_odd m).elim fun hm ↦ by rw [even_iff.1 hm, odd_iff.1 (hm.add_odd hn)]) fun hm ↦ by diff --git a/Mathlib/Combinatorics/SetFamily/KruskalKatona.lean b/Mathlib/Combinatorics/SetFamily/KruskalKatona.lean index 29713f6399feb..8b5ce0135fec7 100644 --- a/Mathlib/Combinatorics/SetFamily/KruskalKatona.lean +++ b/Mathlib/Combinatorics/SetFamily/KruskalKatona.lean @@ -52,10 +52,6 @@ namespace Finset namespace Colex variable {α : Type*} [LinearOrder α] {𝒜 𝒜₁ 𝒜₂ : Finset (Finset α)} {s t : Finset α} {r : ℕ} -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ /-- This is important for iterating Kruskal-Katona: the shadow of an initial segment is also an initial segment. -/ lemma shadow_initSeg [Fintype α] (hs : s.Nonempty) : @@ -71,7 +67,7 @@ lemma shadow_initSeg [Fintype α] (hs : s.Nonempty) : · simpa [ha] using erase_le_erase_min' hts hst.ge (mem_insert_self _ _) -- Now show that if t ≤ s - min s, there is j such that t ∪ j ≤ s -- We choose j as the smallest thing not in t - simp_rw [le_iff_eq_or_lt, lt_iff_exists_filter_lt, mem_sdiff, filter_inj, _root_.and_assoc] + simp_rw [le_iff_eq_or_lt, lt_iff_exists_filter_lt, mem_sdiff, filter_inj, and_assoc] simp only [toColex_inj, ofColex_toColex, ne_eq, and_imp] rintro cards' (rfl | ⟨k, hks, hkt, z⟩) -- If t = s - min s, then use j = min s so t ∪ j = s @@ -129,17 +125,13 @@ variable {α : Type*} [LinearOrder α] {s U V : Finset α} {n : ℕ} namespace UV -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ /-- Applying the compression makes the set smaller in colex. This is intuitive since a portion of the set is being "shifted down" as `max U < max V`. -/ lemma toColex_compress_lt_toColex {hU : U.Nonempty} {hV : V.Nonempty} (h : max' U hU < max' V hV) (hA : compress U V s ≠ s) : toColex (compress U V s) < toColex s := by rw [compress, ite_ne_right_iff] at hA rw [compress, if_pos hA.1, lt_iff_exists_filter_lt] - simp_rw [mem_sdiff (s := s), filter_inj, _root_.and_assoc] + simp_rw [mem_sdiff (s := s), filter_inj, and_assoc] refine ⟨_, hA.1.2 <| max'_mem _ hV, not_mem_sdiff_of_mem_right <| max'_mem _ _, fun a ha ↦ ?_⟩ have : a ∉ V := fun H ↦ ha.not_le (le_max' _ _ H) have : a ∉ U := fun H ↦ ha.not_lt ((le_max' _ _ H).trans_lt h) diff --git a/Mathlib/Combinatorics/SimpleGraph/Regularity/Equitabilise.lean b/Mathlib/Combinatorics/SimpleGraph/Regularity/Equitabilise.lean index e74fad0ddff83..96d140587f9cc 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Regularity/Equitabilise.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Regularity/Equitabilise.lean @@ -32,10 +32,6 @@ namespace Finpartition variable {α : Type*} [DecidableEq α] {s t : Finset α} {m n a b : ℕ} {P : Finpartition s} -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ /-- Given a partition `P` of `s`, as well as a proof that `a * m + b * (m + 1) = s.card`, we can find a new partition `Q` of `s` where each part has size `m` or `m + 1`, every part of `P` is the union of parts of `Q` plus at most `m` extra elements, there are `b` parts of size `m + 1` and @@ -50,7 +46,7 @@ theorem equitabilise_aux (hs : a * m + b * (m + 1) = s.card) : obtain rfl | m_pos := m.eq_zero_or_pos · refine ⟨⊥, by simp, ?_, by simpa [Finset.filter_true_of_mem] using hs.symm⟩ simp only [le_zero_iff, card_eq_zero, mem_biUnion, exists_prop, mem_filter, id, - _root_.and_assoc, sdiff_eq_empty_iff_subset, subset_iff] + and_assoc, sdiff_eq_empty_iff_subset, subset_iff] exact fun x hx a ha => ⟨{a}, mem_map_of_mem _ (P.le hx ha), singleton_subset_iff.2 ha, mem_singleton_self _⟩ -- Prove the case `m > 0` by strong induction on `s` diff --git a/Mathlib/Combinatorics/SimpleGraph/Triangle/Basic.lean b/Mathlib/Combinatorics/SimpleGraph/Triangle/Basic.lean index a96ef7c723fb8..c3c8b49e6cd43 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Triangle/Basic.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Triangle/Basic.lean @@ -83,10 +83,6 @@ lemma LocallyLinear.map (f : α ↪ β) (hG : G.LocallyLinear) : (G.map f).Local · rw [← Equiv.coe_toEmbedding, ← map_symm] exact LocallyLinear.map _ -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ lemma edgeDisjointTriangles_iff_mem_sym2_subsingleton : G.EdgeDisjointTriangles ↔ ∀ ⦃e : Sym2 α⦄, ¬ e.IsDiag → {s ∈ G.cliqueSet 3 | e ∈ (s : Finset α).sym2}.Subsingleton := by @@ -96,7 +92,7 @@ lemma edgeDisjointTriangles_iff_mem_sym2_subsingleton : ext s simp only [mem_sym2_iff, Sym2.mem_iff, forall_eq_or_imp, forall_eq, Set.sep_and, Set.mem_inter_iff, Set.mem_sep_iff, mem_cliqueSet_iff, Set.mem_setOf_eq, - and_and_and_comm (b := _ ∈ _), _root_.and_self, is3Clique_iff] + and_and_and_comm (b := _ ∈ _), and_self, is3Clique_iff] constructor · rintro ⟨⟨c, d, e, hcd, hce, hde, rfl⟩, hab⟩ simp only [mem_insert, mem_singleton] at hab diff --git a/Mathlib/Data/Finset/NatDivisors.lean b/Mathlib/Data/Finset/NatDivisors.lean index 03edf844dbd26..cbc7aa896c2f5 100644 --- a/Mathlib/Data/Finset/NatDivisors.lean +++ b/Mathlib/Data/Finset/NatDivisors.lean @@ -16,16 +16,12 @@ exhibiting `Nat.divisors` as a multiplicative homomorphism from `ℕ` to `Finset open Nat Finset open scoped Pointwise -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ /-- The divisors of a product of natural numbers are the pointwise product of the divisors of the factors. -/ lemma Nat.divisors_mul (m n : ℕ) : divisors (m * n) = divisors m * divisors n := by ext k simp_rw [mem_mul, mem_divisors, dvd_mul, mul_ne_zero_iff, ← exists_and_left, ← exists_and_right] - simp only [_root_.and_assoc, _root_.and_comm, and_left_comm] + simp only [and_assoc, and_comm, and_left_comm] /-- `Nat.divisors` as a `MonoidHom`. -/ @[simps] diff --git a/Mathlib/Data/Int/Defs.lean b/Mathlib/Data/Int/Defs.lean index 5312b39a07639..6bf06c9ff023a 100644 --- a/Mathlib/Data/Int/Defs.lean +++ b/Mathlib/Data/Int/Defs.lean @@ -43,12 +43,7 @@ protected lemma le_antisymm_iff : a = b ↔ a ≤ b ∧ b ≤ a := protected lemma le_iff_eq_or_lt : a ≤ b ↔ a = b ∨ a < b := by rw [Int.le_antisymm_iff, Int.lt_iff_le_not_le, ← and_or_left]; simp [em] -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ -protected lemma le_iff_lt_or_eq : a ≤ b ↔ a < b ∨ a = b := by rw [Int.le_iff_eq_or_lt, - _root_.or_comm] +protected lemma le_iff_lt_or_eq : a ≤ b ↔ a < b ∨ a = b := by rw [Int.le_iff_eq_or_lt, or_comm] end Order diff --git a/Mathlib/Data/List/Basic.lean b/Mathlib/Data/List/Basic.lean index 43008b218c99b..d3d29ac80d33d 100644 --- a/Mathlib/Data/List/Basic.lean +++ b/Mathlib/Data/List/Basic.lean @@ -2177,18 +2177,13 @@ end Forall theorem get_attach (L : List α) (i) : (L.attach.get i).1 = L.get ⟨i, length_attach (L := L) ▸ i.2⟩ := by simp -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ @[simp 1100] theorem mem_map_swap (x : α) (y : β) (xs : List (α × β)) : (y, x) ∈ map Prod.swap xs ↔ (x, y) ∈ xs := by induction' xs with x xs xs_ih · simp only [not_mem_nil, map_nil] · cases' x with a b - simp only [mem_cons, Prod.mk.inj_iff, map, Prod.swap_prod_mk, Prod.exists, xs_ih, - _root_.and_comm] + simp only [mem_cons, Prod.mk.inj_iff, map, Prod.swap_prod_mk, Prod.exists, xs_ih, and_comm] theorem dropSlice_eq (xs : List α) (n m : ℕ) : dropSlice n m xs = xs.take n ++ xs.drop (n + m) := by induction n generalizing xs diff --git a/Mathlib/Data/List/Chain.lean b/Mathlib/Data/List/Chain.lean index a627cab57f2a6..965aee596b838 100644 --- a/Mathlib/Data/List/Chain.lean +++ b/Mathlib/Data/List/Chain.lean @@ -47,10 +47,6 @@ theorem Chain.iff_mem {a : α} {l : List α} : theorem chain_singleton {a b : α} : Chain R a [b] ↔ R a b := by simp only [chain_cons, Chain.nil, and_true] -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ theorem chain_split {a b : α} {l₁ l₂ : List α} : Chain R a (l₁ ++ b :: l₂) ↔ Chain R a (l₁ ++ [b]) ∧ Chain R b l₂ := by induction' l₁ with x l₁ IH generalizing a <;> @@ -243,18 +239,13 @@ theorem Chain'.cons' {x} : ∀ {l : List α}, Chain' R l → (∀ y ∈ l.head?, theorem chain'_cons' {x l} : Chain' R (x :: l) ↔ (∀ y ∈ head? l, R x y) ∧ Chain' R l := ⟨fun h => ⟨h.rel_head?, h.tail⟩, fun ⟨h₁, h₂⟩ => h₂.cons' h₁⟩ -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefixes below. --/ theorem chain'_append : ∀ {l₁ l₂ : List α}, Chain' R (l₁ ++ l₂) ↔ Chain' R l₁ ∧ Chain' R l₂ ∧ ∀ x ∈ l₁.getLast?, ∀ y ∈ l₂.head?, R x y | [], l => by simp - | [a], l => by simp [chain'_cons', _root_.and_comm] + | [a], l => by simp [chain'_cons', and_comm] | a :: b :: l₁, l₂ => by - rw [cons_append, cons_append, chain'_cons, chain'_cons, ← cons_append, chain'_append, - _root_.and_assoc] + rw [cons_append, cons_append, chain'_cons, chain'_cons, ← cons_append, chain'_append, and_assoc] simp theorem Chain'.append (h₁ : Chain' R l₁) (h₂ : Chain' R l₂) @@ -293,16 +284,12 @@ theorem Chain'.imp_head {x y} (h : ∀ {z}, R x z → R y z) {l} (hl : Chain' R Chain' R (y :: l) := hl.tail.cons' fun _ hz => h <| hl.rel_head? hz -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ theorem chain'_reverse : ∀ {l}, Chain' R (reverse l) ↔ Chain' (flip R) l | [] => Iff.rfl | [a] => by simp only [chain'_singleton, reverse_singleton] | a :: b :: l => by rw [chain'_cons, reverse_cons, reverse_cons, append_assoc, cons_append, nil_append, - chain'_split, ← reverse_cons, @chain'_reverse (b :: l), _root_.and_comm, chain'_pair, flip] + chain'_split, ← reverse_cons, @chain'_reverse (b :: l), and_comm, chain'_pair, flip] theorem chain'_iff_get {R} : ∀ {l : List α}, Chain' R l ↔ ∀ (i : ℕ) (h : i < length l - 1), @@ -320,10 +307,6 @@ theorem Chain'.append_overlap {l₁ l₂ l₃ : List α} (h₁ : Chain' R (l₁ h₁.append h₂.right_of_append <| by simpa only [getLast?_append_of_ne_nil _ hn] using (chain'_append.1 h₂).2.2 -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ lemma chain'_join : ∀ {L : List (List α)}, [] ∉ L → (Chain' R L.join ↔ (∀ l ∈ L, Chain' R l) ∧ L.Chain' (fun l₁ l₂ => ∀ᵉ (x ∈ l₁.getLast?) (y ∈ l₂.head?), R x y)) @@ -333,7 +316,7 @@ lemma chain'_join : ∀ {L : List (List α)}, [] ∉ L → rw [mem_cons, not_or, ← Ne] at hL rw [join, chain'_append, chain'_join hL.2, forall_mem_cons, chain'_cons] rw [mem_cons, not_or, ← Ne] at hL - simp only [forall_mem_cons, _root_.and_assoc, join, head?_append_of_ne_nil _ hL.2.1.symm] + simp only [forall_mem_cons, and_assoc, join, head?_append_of_ne_nil _ hL.2.1.symm] exact Iff.rfl.and (Iff.rfl.and <| Iff.rfl.and and_comm) /-- If `a` and `b` are related by the reflexive transitive closure of `r`, then there is an diff --git a/Mathlib/Data/List/Nodup.lean b/Mathlib/Data/List/Nodup.lean index e8c9adf26282f..f1a4e94793afc 100644 --- a/Mathlib/Data/List/Nodup.lean +++ b/Mathlib/Data/List/Nodup.lean @@ -102,17 +102,13 @@ theorem nodup_iff_get?_ne_get? {l : List α} : l.Nodup ↔ ∀ i j : ℕ, i < j → j < l.length → l.get? i ≠ l.get? j := by simp [nodup_iff_getElem?_ne_getElem?] -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ theorem Nodup.ne_singleton_iff {l : List α} (h : Nodup l) (x : α) : l ≠ [x] ↔ l = [] ∨ ∃ y ∈ l, y ≠ x := by induction' l with hd tl hl · simp · specialize hl h.of_cons by_cases hx : tl = [x] - · simpa [hx, _root_.and_comm, and_or_left] using h + · simpa [hx, and_comm, and_or_left] using h · rw [← Ne, hl] at hx rcases hx with (rfl | ⟨y, hy, hx⟩) · simp @@ -178,13 +174,9 @@ theorem Nodup.append (d₁ : Nodup l₁) (d₂ : Nodup l₂) (dj : Disjoint l₁ theorem nodup_append_comm {l₁ l₂ : List α} : Nodup (l₁ ++ l₂) ↔ Nodup (l₂ ++ l₁) := by simp only [nodup_append, and_left_comm, disjoint_comm] -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ theorem nodup_middle {a : α} {l₁ l₂ : List α} : Nodup (l₁ ++ a :: l₂) ↔ Nodup (a :: (l₁ ++ l₂)) := by - simp only [nodup_append, not_or, and_left_comm, _root_.and_assoc, nodup_cons, mem_append, + simp only [nodup_append, not_or, and_left_comm, and_assoc, nodup_cons, mem_append, disjoint_cons_right] theorem Nodup.of_map (f : α → β) {l : List α} : Nodup (map f l) → Nodup l := @@ -267,14 +259,10 @@ theorem nodup_join {L : List (List α)} : Nodup (join L) ↔ (∀ l ∈ L, Nodup l) ∧ Pairwise Disjoint L := by simp only [Nodup, pairwise_join, disjoint_left.symm, forall_mem_ne] -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ theorem nodup_bind {l₁ : List α} {f : α → List β} : Nodup (l₁.bind f) ↔ (∀ x ∈ l₁, Nodup (f x)) ∧ Pairwise (fun a b : α => Disjoint (f a) (f b)) l₁ := by - simp only [List.bind, nodup_join, pairwise_map, _root_.and_comm, and_left_comm, mem_map, + simp only [List.bind, nodup_join, pairwise_map, and_comm, and_left_comm, mem_map, exists_imp, and_imp] rw [show (∀ (l : List β) (x : α), f x = l → x ∈ l₁ → Nodup l) ↔ ∀ x : α, x ∈ l₁ → Nodup (f x) from forall_swap.trans <| forall_congr' fun _ => forall_eq'] diff --git a/Mathlib/Data/List/Perm.lean b/Mathlib/Data/List/Perm.lean index f0e74dcc0d2b3..14611136f55f6 100644 --- a/Mathlib/Data/List/Perm.lean +++ b/Mathlib/Data/List/Perm.lean @@ -246,17 +246,13 @@ theorem Perm.bagInter {l₁ l₂ t₁ t₂ : List α} (hl : l₁ ~ l₂) (ht : t l₁.bagInter t₁ ~ l₂.bagInter t₂ := ht.bagInter_left l₂ ▸ hl.bagInter_right _ -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ theorem perm_replicate_append_replicate {l : List α} {a b : α} {m n : ℕ} (h : a ≠ b) : l ~ replicate m a ++ replicate n b ↔ count a l = m ∧ count b l = n ∧ l ⊆ [a, b] := by rw [perm_iff_count, ← Decidable.and_forall_ne a, ← Decidable.and_forall_ne b] suffices l ⊆ [a, b] ↔ ∀ c, c ≠ b → c ≠ a → c ∉ l by simp (config := { contextual := true }) [count_replicate, h, this, count_eq_zero, Ne.symm] trans ∀ c, c ∈ l → c = b ∨ c = a - · simp [subset_def, _root_.or_comm] + · simp [subset_def, or_comm] · exact forall_congr' fun _ => by rw [← and_imp, ← not_or, not_imp_not] theorem Perm.dedup {l₁ l₂ : List α} (p : l₁ ~ l₂) : dedup l₁ ~ dedup l₂ := diff --git a/Mathlib/Data/List/Permutation.lean b/Mathlib/Data/List/Permutation.lean index 37fb13a041412..dfa4674956933 100644 --- a/Mathlib/Data/List/Permutation.lean +++ b/Mathlib/Data/List/Permutation.lean @@ -166,10 +166,6 @@ theorem foldr_permutationsAux2 (t : α) (ts : List α) (r L : List (List α)) : · rfl · simp_rw [foldr_cons, ih, bind_cons, append_assoc, permutationsAux2_append] -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ theorem mem_foldr_permutationsAux2 {t : α} {ts : List α} {r L : List (List α)} {l' : List α} : l' ∈ foldr (fun y r => (permutationsAux2 t ts r y id).2) r L ↔ l' ∈ r ∨ ∃ l₁ l₂, l₁ ++ l₂ ∈ L ∧ l₂ ≠ [] ∧ l' = l₁ ++ t :: l₂ ++ ts := by @@ -180,7 +176,7 @@ theorem mem_foldr_permutationsAux2 {t : α} {ts : List α} {r L : List (List α) ⟨fun ⟨_, aL, l₁, l₂, l0, e, h⟩ => ⟨l₁, l₂, l0, e ▸ aL, h⟩, fun ⟨l₁, l₂, l0, aL, h⟩ => ⟨_, aL, l₁, l₂, l0, rfl, h⟩⟩ rw [foldr_permutationsAux2] - simp only [mem_permutationsAux2', ← this, _root_.or_comm, and_left_comm, mem_append, mem_bind, + simp only [mem_permutationsAux2', ← this, or_comm, and_left_comm, mem_append, mem_bind, append_assoc, cons_append, exists_prop] theorem length_foldr_permutationsAux2 (t : α) (ts : List α) (r L : List (List α)) : diff --git a/Mathlib/Data/List/Rotate.lean b/Mathlib/Data/List/Rotate.lean index 923a8ba69b0ed..04072b6a698df 100644 --- a/Mathlib/Data/List/Rotate.lean +++ b/Mathlib/Data/List/Rotate.lean @@ -104,15 +104,11 @@ theorem rotate_cons_succ (l : List α) (a : α) (n : ℕ) : (a :: l : List α).rotate (n + 1) = (l ++ [a]).rotate n := by rw [rotate_eq_rotate', rotate_eq_rotate', rotate'_cons_succ] -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ @[simp] theorem mem_rotate : ∀ {l : List α} {a : α} {n : ℕ}, a ∈ l.rotate n ↔ a ∈ l | [], _, n => by simp | a :: l, _, 0 => by simp - | a :: l, _, n + 1 => by simp [rotate_cons_succ, mem_rotate, _root_.or_comm] + | a :: l, _, n + 1 => by simp [rotate_cons_succ, mem_rotate, or_comm] @[simp] theorem length_rotate (l : List α) (n : ℕ) : (l.rotate n).length = l.length := by diff --git a/Mathlib/Data/Multiset/Basic.lean b/Mathlib/Data/Multiset/Basic.lean index 0ecc3f20a0bed..0c4fb024e449e 100644 --- a/Mathlib/Data/Multiset/Basic.lean +++ b/Mathlib/Data/Multiset/Basic.lean @@ -1832,25 +1832,17 @@ theorem filter_filter (q) [DecidablePred q] (s : Multiset α) : filter p (filter q s) = filter (fun a => p a ∧ q a) s := Quot.inductionOn s fun l => by simp -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ lemma filter_comm (q) [DecidablePred q] (s : Multiset α) : - filter p (filter q s) = filter q (filter p s) := by simp [_root_.and_comm] + filter p (filter q s) = filter q (filter p s) := by simp [and_comm] theorem filter_add_filter (q) [DecidablePred q] (s : Multiset α) : filter p s + filter q s = filter (fun a => p a ∨ q a) s + filter (fun a => p a ∧ q a) s := Multiset.induction_on s rfl fun a s IH => by by_cases p a <;> by_cases q a <;> simp [*] -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ 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 [Decidable.em, -Bool.not_eq_true, -not_and, not_and_or, _root_.or_comm] + · 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] @@ -2690,13 +2682,9 @@ lemma add_eq_union_left_of_le [DecidableEq α] {s t u : Multiset α} (h : t ≤ · rintro ⟨h0, rfl⟩ exact h0 -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ lemma add_eq_union_right_of_le [DecidableEq α] {x y z : Multiset α} (h : z ≤ y) : x + y = x ∪ z ↔ y = z ∧ x.Disjoint y := by - simpa only [_root_.and_comm] using add_eq_union_left_of_le h + simpa only [and_comm] using add_eq_union_left_of_le h theorem disjoint_map_map {f : α → γ} {g : β → γ} {s : Multiset α} {t : Multiset β} : Disjoint (s.map f) (t.map g) ↔ ∀ a ∈ s, ∀ b ∈ t, f a ≠ g b := by diff --git a/Mathlib/Data/Nat/Factorization/Basic.lean b/Mathlib/Data/Nat/Factorization/Basic.lean index 1285453521a3b..27ddcf0d3a424 100644 --- a/Mathlib/Data/Nat/Factorization/Basic.lean +++ b/Mathlib/Data/Nat/Factorization/Basic.lean @@ -455,10 +455,6 @@ theorem setOf_pow_dvd_eq_Icc_factorization {n p : ℕ} (pp : p.Prime) (hn : n ext simp [Nat.lt_succ_iff, one_le_iff_ne_zero, pp.pow_dvd_iff_le_factorization hn] -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ /-- The set of positive powers of prime `p` that divide `n` is exactly the set of positive natural numbers up to `n.factorization p`. -/ theorem Icc_factorization_eq_pow_dvd (n : ℕ) {p : ℕ} (pp : Prime p) : @@ -466,7 +462,7 @@ theorem Icc_factorization_eq_pow_dvd (n : ℕ) {p : ℕ} (pp : Prime p) : rcases eq_or_ne n 0 with (rfl | hn) · simp ext x - simp only [mem_Icc, Finset.mem_filter, mem_Ico, _root_.and_assoc, and_congr_right_iff, + simp only [mem_Icc, Finset.mem_filter, mem_Ico, and_assoc, and_congr_right_iff, pp.pow_dvd_iff_le_factorization hn, iff_and_self] exact fun _ H => lt_of_le_of_lt H (factorization_lt p hn) diff --git a/Mathlib/Data/Nat/Log.lean b/Mathlib/Data/Nat/Log.lean index 744b67ca07c10..f2cf3df6eca07 100644 --- a/Mathlib/Data/Nat/Log.lean +++ b/Mathlib/Data/Nat/Log.lean @@ -121,15 +121,11 @@ lemma log_le_self (b x : ℕ) : log b x ≤ x := theorem lt_pow_succ_log_self {b : ℕ} (hb : 1 < b) (x : ℕ) : x < b ^ (log b x).succ := lt_pow_of_log_lt hb (lt_succ_self _) -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ theorem log_eq_iff {b m n : ℕ} (h : m ≠ 0 ∨ 1 < b ∧ n ≠ 0) : log b n = m ↔ b ^ m ≤ n ∧ n < b ^ (m + 1) := by rcases em (1 < b ∧ n ≠ 0) with (⟨hb, hn⟩ | hbn) · rw [le_antisymm_iff, ← Nat.lt_succ_iff, ← pow_le_iff_le_log, ← lt_pow_iff_log_lt, - _root_.and_comm] <;> assumption + and_comm] <;> assumption have hm : m ≠ 0 := h.resolve_right hbn rw [not_and_or, not_lt, Ne, not_not] at hbn rcases hbn with (hb | rfl) diff --git a/Mathlib/Data/Nat/WithBot.lean b/Mathlib/Data/Nat/WithBot.lean index 418d1347bbc66..125eeae18ef2b 100644 --- a/Mathlib/Data/Nat/WithBot.lean +++ b/Mathlib/Data/Nat/WithBot.lean @@ -30,13 +30,9 @@ theorem add_eq_zero_iff {n m : WithBot ℕ} : n + m = 0 ↔ n = 0 ∧ m = 0 := b · simp [WithBot.add_bot] simp [← WithBot.coe_add, add_eq_zero_iff_of_nonneg] -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ theorem add_eq_one_iff {n m : WithBot ℕ} : n + m = 1 ↔ n = 0 ∧ m = 1 ∨ n = 1 ∧ m = 0 := by cases n - · simp only [WithBot.bot_add, WithBot.bot_ne_one, WithBot.bot_ne_zero, false_and, _root_.or_self] + · simp only [WithBot.bot_add, WithBot.bot_ne_one, WithBot.bot_ne_zero, false_and, or_self] cases m · simp [WithBot.add_bot] simp [← WithBot.coe_add, Nat.add_eq_one_iff] diff --git a/Mathlib/Data/PFunctor/Univariate/M.lean b/Mathlib/Data/PFunctor/Univariate/M.lean index 699b8b76a816f..73fe251a4dda5 100644 --- a/Mathlib/Data/PFunctor/Univariate/M.lean +++ b/Mathlib/Data/PFunctor/Univariate/M.lean @@ -470,10 +470,6 @@ theorem corec_def {X} (f : X → F X) (x₀ : X) : M.corec f x₀ = M.mk (F.map dsimp only [PFunctor.map] congr -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ theorem ext_aux [Inhabited (M F)] [DecidableEq F.A] {n : ℕ} (x y z : M F) (hx : Agree' n z x) (hy : Agree' n z y) (hrec : ∀ ps : Path F, n = ps.length → iselect ps x = iselect ps y) : x.approx (n + 1) = y.approx (n + 1) := by diff --git a/Mathlib/Logic/Encodable/Basic.lean b/Mathlib/Logic/Encodable/Basic.lean index c111c09a26737..60412ec201292 100644 --- a/Mathlib/Logic/Encodable/Basic.lean +++ b/Mathlib/Logic/Encodable/Basic.lean @@ -555,13 +555,9 @@ theorem sequence_mono_nat {r : β → β → Prop} {f : α → β} (hf : Directe · exact (Classical.choose_spec (hf p p)).1 · exact (Classical.choose_spec (hf p a)).1 -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ theorem rel_sequence {r : β → β → Prop} {f : α → β} (hf : Directed r f) (a : α) : r (f a) (f (hf.sequence f (encode a + 1))) := by - simp only [Directed.sequence, add_eq, add_zero, encodek, _root_.and_self] + simp only [Directed.sequence, add_eq, add_zero, encodek, and_self] exact (Classical.choose_spec (hf _ a)).2 variable [Preorder β] {f : α → β} diff --git a/Mathlib/NumberTheory/Bertrand.lean b/Mathlib/NumberTheory/Bertrand.lean index 7f3a1d190e6f7..772c5d12f2dff 100644 --- a/Mathlib/NumberTheory/Bertrand.lean +++ b/Mathlib/NumberTheory/Bertrand.lean @@ -121,10 +121,6 @@ theorem bertrand_main_inequality {n : ℕ} (n_large : 512 ≤ n) : · norm_num1 · exact cast_div_le.trans (by norm_cast) -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ /-- A lemma that tells us that, in the case where Bertrand's postulate does not hold, the prime factorization of the central binomial coefficient only has factors at most `2 * n / 3 + 1`. -/ @@ -138,7 +134,7 @@ theorem centralBinom_factorization_small (n : ℕ) (n_large : 2 < n) rw [Finset.mem_range, Nat.lt_succ_iff] at hx h2x rw [not_le, div_lt_iff_lt_mul' three_pos, mul_comm x] at h2x replace no_prime := not_exists.mp no_prime x - rw [← _root_.and_assoc, not_and', not_and_or, not_lt] at no_prime + rw [← and_assoc, not_and', not_and_or, not_lt] at no_prime cases' no_prime hx with h h · rw [factorization_eq_zero_of_non_prime n.centralBinom h, Nat.pow_zero] · rw [factorization_centralBinom_of_two_mul_self_lt_three_mul n_large h h2x, Nat.pow_zero] diff --git a/Mathlib/NumberTheory/Divisors.lean b/Mathlib/NumberTheory/Divisors.lean index 43de39c1d54b0..0c98cefad6ecd 100644 --- a/Mathlib/NumberTheory/Divisors.lean +++ b/Mathlib/NumberTheory/Divisors.lean @@ -66,14 +66,10 @@ theorem filter_dvd_eq_properDivisors (h : n ≠ 0) : theorem properDivisors.not_self_mem : ¬n ∈ properDivisors n := by simp [properDivisors] -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ @[simp] theorem mem_properDivisors {m : ℕ} : n ∈ properDivisors m ↔ n ∣ m ∧ n < m := by rcases eq_or_ne m 0 with (rfl | hm); · simp [properDivisors] - simp only [_root_.and_comm, ← filter_dvd_eq_properDivisors hm, mem_filter, mem_range] + simp only [and_comm, ← filter_dvd_eq_properDivisors hm, mem_filter, mem_range] theorem insert_self_properDivisors (h : n ≠ 0) : insert n (properDivisors n) = divisors n := by rw [divisors, properDivisors, Ico_succ_right_eq_insert_Ico (one_le_iff_ne_zero.2 h), @@ -100,15 +96,11 @@ theorem dvd_of_mem_divisors {m : ℕ} (h : n ∈ divisors m) : n ∣ m := by · apply dvd_zero · simp [mem_divisors.1 h] -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ @[simp] theorem mem_divisorsAntidiagonal {x : ℕ × ℕ} : x ∈ divisorsAntidiagonal n ↔ x.fst * x.snd = n ∧ n ≠ 0 := by simp only [divisorsAntidiagonal, Finset.mem_Ico, Ne, Finset.mem_filter, Finset.mem_product] - rw [_root_.and_comm] + rw [and_comm] apply and_congr_right rintro rfl constructor <;> intro h @@ -429,14 +421,10 @@ theorem sum_properDivisors_eq_one_iff_prime : ∑ x ∈ n.properDivisors, x = 1 (one_mem_properDivisors_iff_one_lt.2 (succ_lt_succ (Nat.succ_pos _)))) ((sum_singleton _ _).trans h.symm) -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ theorem mem_properDivisors_prime_pow {p : ℕ} (pp : p.Prime) (k : ℕ) {x : ℕ} : x ∈ properDivisors (p ^ k) ↔ ∃ (j : ℕ) (_ : j < k), x = p ^ j := by rw [mem_properDivisors, Nat.dvd_prime_pow pp, ← exists_and_right] - simp only [exists_prop, _root_.and_assoc] + simp only [exists_prop, and_assoc] apply exists_congr intro a constructor <;> intro h @@ -481,17 +469,13 @@ theorem prod_divisorsAntidiagonal' {M : Type*} [CommMonoid M] (f : ℕ → ℕ rw [← map_swap_divisorsAntidiagonal, Finset.prod_map] exact prod_divisorsAntidiagonal fun i j => f j i -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ /-- The factors of `n` are the prime divisors -/ theorem primeFactors_eq_to_filter_divisors_prime (n : ℕ) : n.primeFactors = (divisors n).filter Prime := by rcases n.eq_zero_or_pos with (rfl | hn) · simp · ext q - simpa [hn, hn.ne', mem_primeFactorsList] using _root_.and_comm + simpa [hn, hn.ne', mem_primeFactorsList] using and_comm @[deprecated (since := "2024-07-17")] alias prime_divisors_eq_to_filter_divisors_prime := primeFactors_eq_to_filter_divisors_prime diff --git a/Mathlib/NumberTheory/LegendreSymbol/GaussEisensteinLemmas.lean b/Mathlib/NumberTheory/LegendreSymbol/GaussEisensteinLemmas.lean index b2cb30daa0d3e..392736e26084d 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/GaussEisensteinLemmas.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/GaussEisensteinLemmas.lean @@ -180,10 +180,6 @@ private theorem sum_Ico_eq_card_lt {p q : ℕ} : (by simp (config := { contextual := true }) only [mem_filter, mem_sigma, and_self_iff, forall_true_iff, mem_product]) (fun _ _ ↦ rfl) (fun _ _ ↦ rfl) -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ /-- Each of the sums in this lemma is the cardinality of the set of integer points in each of the two triangles formed by the diagonal of the rectangle `(0, p/2) × (0, q/2)`. Adding them gives the number of points in the rectangle. -/ @@ -197,7 +193,7 @@ theorem sum_mul_div_add_sum_mul_div_eq_mul (p q : ℕ) [hp : Fact p.Prime] (hq0 card_equiv (Equiv.prodComm _ _) (fun ⟨_, _⟩ => by simp (config := { contextual := true }) only [mem_filter, and_self_iff, Prod.swap_prod_mk, - forall_true_iff, mem_product, Equiv.prodComm_apply, _root_.and_assoc, and_left_comm]) + forall_true_iff, mem_product, Equiv.prodComm_apply, and_assoc, and_left_comm]) have hdisj : Disjoint ((Ico 1 (p / 2).succ ×ˢ Ico 1 (q / 2).succ).filter fun x : ℕ × ℕ => x.2 * p ≤ x.1 * q) ((Ico 1 (p / 2).succ ×ˢ Ico 1 (q / 2).succ).filter fun x : ℕ × ℕ => x.1 * q ≤ x.2 * p) := by diff --git a/Mathlib/NumberTheory/SumFourSquares.lean b/Mathlib/NumberTheory/SumFourSquares.lean index a08b8c479c8ee..5d7a765adee86 100644 --- a/Mathlib/NumberTheory/SumFourSquares.lean +++ b/Mathlib/NumberTheory/SumFourSquares.lean @@ -121,10 +121,6 @@ private theorem sum_four_squares_of_two_mul_sum_four_squares {m a b c d : ℤ} have : (∑ x, f (σ x) ^ 2) = ∑ x, f x ^ 2 := Equiv.sum_comp σ (f · ^ 2) simpa only [← hx, ← hy, Fin.sum_univ_four, add_assoc] using this -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ /-- Lagrange's **four squares theorem** for a prime number. Use `Nat.sum_four_squares` instead. -/ protected theorem Prime.sum_four_squares {p : ℕ} (hp : p.Prime) : ∃ a b c d : ℕ, a ^ 2 + b ^ 2 + c ^ 2 + d ^ 2 = p := by @@ -174,10 +170,9 @@ protected theorem Prime.sum_four_squares {p : ℕ} (hp : p.Prime) : -- The quotient `r` is not zero, because otherwise `f a = f b = f c = f d = 0`, hence -- `m` divides each `a`, `b`, `c`, `d`, thus `m ∣ p` which is impossible. rcases (zero_le r).eq_or_gt with rfl | hr₀ - · replace hr : f a = 0 ∧ f b = 0 ∧ f c = 0 ∧ f d = 0 := by simpa [_root_.and_assoc] using hr + · replace hr : f a = 0 ∧ f b = 0 ∧ f c = 0 ∧ f d = 0 := by simpa [and_assoc] using hr obtain ⟨⟨a, rfl⟩, ⟨b, rfl⟩, ⟨c, rfl⟩, ⟨d, rfl⟩⟩ : m ∣ a ∧ m ∣ b ∧ m ∣ c ∧ m ∣ d := by - simp only [← ZMod.natCast_zmod_eq_zero_iff_dvd, ← hf_mod, hr, Int.cast_zero, - _root_.and_self] + simp only [← ZMod.natCast_zmod_eq_zero_iff_dvd, ← hf_mod, hr, Int.cast_zero, and_self] have : m * m ∣ m * p := habcd ▸ ⟨a ^ 2 + b ^ 2 + c ^ 2 + d ^ 2, by ring⟩ rw [mul_dvd_mul_iff_left hm₀.ne'] at this exact (hp.eq_one_or_self_of_dvd _ this).elim hm₁.ne' hmp.ne diff --git a/Mathlib/Order/Interval/Finset/Nat.lean b/Mathlib/Order/Interval/Finset/Nat.lean index d59c44eaa2139..8be56194d6481 100644 --- a/Mathlib/Order/Interval/Finset/Nat.lean +++ b/Mathlib/Order/Interval/Finset/Nat.lean @@ -185,14 +185,10 @@ theorem Ico_image_const_sub_eq_Ico (hac : a ≤ c) : rintro ⟨x, hx, rfl⟩ omega -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ theorem Ico_succ_left_eq_erase_Ico : Ico a.succ b = erase (Ico a b) a := by ext x - rw [Ico_succ_left, mem_erase, mem_Ico, mem_Ioo, ← _root_.and_assoc, ne_comm, - _root_.and_comm (a := a ≠ x), lt_iff_le_and_ne] + rw [Ico_succ_left, mem_erase, mem_Ico, mem_Ioo, ← and_assoc, ne_comm, + and_comm (a := a ≠ x), lt_iff_le_and_ne] theorem mod_injOn_Ico (n a : ℕ) : Set.InjOn (· % a) (Finset.Ico n (n + a)) := by induction' n with n ih diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric/NewtonIdentities.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric/NewtonIdentities.lean index ddadb79adcf83..3aa09ead7b4ab 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric/NewtonIdentities.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric/NewtonIdentities.lean @@ -250,17 +250,13 @@ theorem sum_antidiagonal_card_esymm_psum_eq_zero : simp [← sum_filter_add_sum_filter_not (antidiagonal k) (fun a ↦ a.fst < k), ← mul_esymm_eq_sum, mul_add, ← mul_assoc, ← pow_add, mul_comm ↑k (esymm σ R k)] -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ /-- A version of Newton's identities which may be more useful in the case that we know the values of the elementary symmetric polynomials and would like to calculate the values of the power sums. -/ theorem psum_eq_mul_esymm_sub_sum (k : ℕ) (h : 0 < k) : psum σ R k = (-1) ^ (k + 1) * k * esymm σ R k - ∑ a ∈ (antidiagonal k).filter (fun a ↦ a.fst ∈ Set.Ioo 0 k), (-1) ^ a.fst * esymm σ R a.fst * psum σ R a.snd := by - simp only [Set.Ioo, Set.mem_setOf_eq, _root_.and_comm] + simp only [Set.Ioo, Set.mem_setOf_eq, and_comm] have hesymm := mul_esymm_eq_sum σ R k rw [← (sum_filter_add_sum_filter_not ((antidiagonal k).filter (fun a ↦ a.fst < k)) (fun a ↦ 0 < a.fst) (fun a ↦ (-1) ^ a.fst * esymm σ R a.fst * psum σ R a.snd))] at hesymm diff --git a/Mathlib/RingTheory/Polynomial/Cyclotomic/Eval.lean b/Mathlib/RingTheory/Polynomial/Cyclotomic/Eval.lean index 43e050ec7489d..4a1ad15664ee8 100644 --- a/Mathlib/RingTheory/Polynomial/Cyclotomic/Eval.lean +++ b/Mathlib/RingTheory/Polynomial/Cyclotomic/Eval.lean @@ -104,17 +104,12 @@ theorem cyclotomic_pos {n : ℕ} (hn : 2 < n) {R} [LinearOrderedCommRing R] (x : exact hn'.ne' hi.2.2.1 · simpa only [eval_X, eval_one, cyclotomic_two, eval_add] using h.right.le -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ theorem cyclotomic_pos_and_nonneg (n : ℕ) {R} [LinearOrderedCommRing R] (x : R) : (1 < x → 0 < eval x (cyclotomic n R)) ∧ (1 ≤ x → 0 ≤ eval x (cyclotomic n R)) := by rcases n with (_ | _ | _ | n) - · simp only [cyclotomic_zero, eval_one, zero_lt_one, implies_true, zero_le_one, - _root_.and_self] + · simp only [cyclotomic_zero, eval_one, zero_lt_one, implies_true, zero_le_one, and_self] · simp only [zero_add, cyclotomic_one, eval_sub, eval_X, eval_one, sub_pos, imp_self, sub_nonneg, - _root_.and_self] + and_self] · simp only [zero_add, reduceAdd, cyclotomic_two, eval_add, eval_X, eval_one] constructor <;> intro <;> linarith · constructor <;> intro <;> [skip; apply le_of_lt] <;> apply cyclotomic_pos (by omega) From 912f72d369210c9e3e8715f4796a50a330d34df7 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Mon, 7 Oct 2024 17:59:10 +0000 Subject: [PATCH 26/43] feat(SetTheory/Ordinal/Exponential): more lemmas on `Ordinal.log` (#15990) We show multiple variants of `opow_le_iff_le_log` with slightly different assumptions. --- Mathlib/SetTheory/Ordinal/Exponential.lean | 74 +++++++++++++++++++--- 1 file changed, 66 insertions(+), 8 deletions(-) diff --git a/Mathlib/SetTheory/Ordinal/Exponential.lean b/Mathlib/SetTheory/Ordinal/Exponential.lean index f5d3ed03ea35f..ca373660f3136 100644 --- a/Mathlib/SetTheory/Ordinal/Exponential.lean +++ b/Mathlib/SetTheory/Ordinal/Exponential.lean @@ -86,6 +86,14 @@ theorem opow_pos {a : Ordinal} (b : Ordinal) (a0 : 0 < a) : 0 < a ^ b := by theorem opow_ne_zero {a : Ordinal} (b : Ordinal) (a0 : a ≠ 0) : a ^ b ≠ 0 := Ordinal.pos_iff_ne_zero.1 <| opow_pos b <| Ordinal.pos_iff_ne_zero.2 a0 +@[simp] +theorem opow_eq_zero {a b : Ordinal} : a ^ b = 0 ↔ a = 0 ∧ b ≠ 0 := by + obtain rfl | ha := eq_or_ne a 0 + · obtain rfl | hb := eq_or_ne b 0 + · simp + · simp [hb] + · simp [opow_ne_zero b ha, ha] + @[simp, norm_cast] theorem opow_natCast (a : Ordinal) (n : ℕ) : a ^ (n : Ordinal) = a ^ n := by induction n with @@ -286,17 +294,67 @@ theorem opow_log_le_self (b : Ordinal) {x : Ordinal} (hx : x ≠ 0) : b ^ log b rwa [← succ_log_def hb hx] at this · rwa [one_opow, one_le_iff_ne_zero] -/-- `opow b` and `log b` (almost) form a Galois connection. -/ -theorem opow_le_iff_le_log {b x c : Ordinal} (hb : 1 < b) (hx : x ≠ 0) : b ^ c ≤ x ↔ c ≤ log b x := - ⟨fun h => - le_of_not_lt fun hn => - (lt_opow_succ_log_self hb x).not_le <| - ((opow_le_opow_iff_right hb).2 (succ_le_of_lt hn)).trans h, - fun h => ((opow_le_opow_iff_right hb).2 h).trans (opow_log_le_self b hx)⟩ - +/-- `opow b` and `log b` (almost) form a Galois connection. + +See `opow_le_iff_le_log'` for a variant assuming `c ≠ 0` rather than `x ≠ 0`. See also +`le_log_of_opow_le` and `opow_le_of_le_log`, which are both separate implications under weaker +assumptions. -/ +theorem opow_le_iff_le_log {b x c : Ordinal} (hb : 1 < b) (hx : x ≠ 0) : + b ^ c ≤ x ↔ c ≤ log b x := by + constructor <;> + intro h + · apply le_of_not_lt + intro hn + apply (lt_opow_succ_log_self hb x).not_le <| + ((opow_le_opow_iff_right hb).2 <| succ_le_of_lt hn).trans h + · exact ((opow_le_opow_iff_right hb).2 h).trans <| opow_log_le_self b hx + +/-- `opow b` and `log b` (almost) form a Galois connection. + +See `opow_le_iff_le_log` for a variant assuming `x ≠ 0` rather than `c ≠ 0`. See also +`le_log_of_opow_le` and `opow_le_of_le_log`, which are both separate implications under weaker +assumptions. -/ +theorem opow_le_iff_le_log' {b x c : Ordinal} (hb : 1 < b) (hc : c ≠ 0) : + b ^ c ≤ x ↔ c ≤ log b x := by + obtain rfl | hx := eq_or_ne x 0 + · rw [log_zero_right, Ordinal.le_zero, Ordinal.le_zero, opow_eq_zero] + simp [hc, (zero_lt_one.trans hb).ne'] + · exact opow_le_iff_le_log hb hx + +theorem le_log_of_opow_le {b x c : Ordinal} (hb : 1 < b) (h : b ^ c ≤ x) : c ≤ log b x := by + obtain rfl | hx := eq_or_ne x 0 + · rw [Ordinal.le_zero, opow_eq_zero] at h + exact (zero_lt_one.asymm <| h.1 ▸ hb).elim + · exact (opow_le_iff_le_log hb hx).1 h + +theorem opow_le_of_le_log {b x c : Ordinal} (hc : c ≠ 0) (h : c ≤ log b x) : b ^ c ≤ x := by + obtain hb | hb := le_or_lt b 1 + · rw [log_of_left_le_one hb] at h + exact (h.not_lt (Ordinal.pos_iff_ne_zero.2 hc)).elim + · rwa [opow_le_iff_le_log' hb hc] + +/-- `opow b` and `log b` (almost) form a Galois connection. + +See `lt_opow_iff_log_lt'` for a variant assuming `c ≠ 0` rather than `x ≠ 0`. See also +`lt_opow_of_log_lt` and `lt_log_of_lt_opow`, which are both separate implications under weaker +assumptions. -/ theorem lt_opow_iff_log_lt {b x c : Ordinal} (hb : 1 < b) (hx : x ≠ 0) : x < b ^ c ↔ log b x < c := lt_iff_lt_of_le_iff_le (opow_le_iff_le_log hb hx) +/-- `opow b` and `log b` (almost) form a Galois connection. + +See `lt_opow_iff_log_lt` for a variant assuming `x ≠ 0` rather than `c ≠ 0`. See also +`lt_opow_of_log_lt` and `lt_log_of_lt_opow`, which are both separate implications under weaker +assumptions. -/ +theorem lt_opow_iff_log_lt' {b x c : Ordinal} (hb : 1 < b) (hc : c ≠ 0) : x < b ^ c ↔ log b x < c := + lt_iff_lt_of_le_iff_le (opow_le_iff_le_log' hb hc) + +theorem lt_opow_of_log_lt {b x c : Ordinal} (hb : 1 < b) : log b x < c → x < b ^ c := + lt_imp_lt_of_le_imp_le <| le_log_of_opow_le hb + +theorem lt_log_of_lt_opow {b x c : Ordinal} (hc : c ≠ 0) : x < b ^ c → log b x < c := + lt_imp_lt_of_le_imp_le <| opow_le_of_le_log hc + theorem log_pos {b o : Ordinal} (hb : 1 < b) (ho : o ≠ 0) (hbo : b ≤ o) : 0 < log b o := by rwa [← succ_le_iff, succ_zero, ← opow_le_iff_le_log hb ho, opow_one] From 8d75be4fa83128a02844a9c91bdbe52034a0816b Mon Sep 17 00:00:00 2001 From: YnirPaz Date: Mon, 7 Oct 2024 17:59:11 +0000 Subject: [PATCH 27/43] refactor(SetTheory/Ordinal/Arithmetic): change `mk_initialSeg` to work with `Set.Iio` (#16929) change `mk_initialSeg` to work with `Set.Iio` and rename it `mk_Iio_ordinal`. --- Mathlib/SetTheory/Ordinal/Arithmetic.lean | 22 ++++++++++++++----- .../Ordinal/FixedPointApproximants.lean | 2 +- 2 files changed, 18 insertions(+), 6 deletions(-) diff --git a/Mathlib/SetTheory/Ordinal/Arithmetic.lean b/Mathlib/SetTheory/Ordinal/Arithmetic.lean index 331d539469d54..edc763c29c234 100644 --- a/Mathlib/SetTheory/Ordinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Ordinal/Arithmetic.lean @@ -361,16 +361,28 @@ theorem bounded_singleton {r : α → α → Prop} [IsWellOrder α r] (hr : (typ rw [@enum_lt_enum _ r, Subtype.mk_lt_mk] apply lt_succ --- Porting note: `· < ·` requires a type ascription for an `IsWellOrder` instance. -theorem type_subrel_lt (o : Ordinal.{u}) : - type (@Subrel Ordinal (· < ·) { o' : Ordinal | o' < o }) = Ordinal.lift.{u + 1} o := by +@[simp] +theorem typein_ordinal (o : Ordinal.{u}) : + @typein Ordinal (· < ·) _ o = Ordinal.lift.{u + 1} o := by refine Quotient.inductionOn o ?_ rintro ⟨α, r, wo⟩; apply Quotient.sound constructor; refine ((RelIso.preimage Equiv.ulift r).trans (enum r).symm).symm +-- Porting note: `· < ·` requires a type ascription for an `IsWellOrder` instance. +@[deprecated typein_ordinal (since := "2024-09-19")] +theorem type_subrel_lt (o : Ordinal.{u}) : + type (@Subrel Ordinal (· < ·) { o' : Ordinal | o' < o }) = Ordinal.lift.{u + 1} o := + typein_ordinal o + +theorem mk_Iio_ordinal (o : Ordinal.{u}) : + #(Iio o) = Cardinal.lift.{u + 1} o.card := by + rw [lift_card, ← typein_ordinal] + rfl + +@[deprecated mk_Iio_ordinal (since := "2024-09-19")] theorem mk_initialSeg (o : Ordinal.{u}) : - #{ o' : Ordinal | o' < o } = Cardinal.lift.{u + 1} o.card := by - rw [lift_card, ← type_subrel_lt, card_type] + #{ o' : Ordinal | o' < o } = Cardinal.lift.{u + 1} o.card := mk_Iio_ordinal o + /-! ### Normal ordinal functions -/ diff --git a/Mathlib/SetTheory/Ordinal/FixedPointApproximants.lean b/Mathlib/SetTheory/Ordinal/FixedPointApproximants.lean index ddec8ac7b3058..97272de665f4f 100644 --- a/Mathlib/SetTheory/Ordinal/FixedPointApproximants.lean +++ b/Mathlib/SetTheory/Ordinal/FixedPointApproximants.lean @@ -50,7 +50,7 @@ theorem not_injective_limitation_set : ¬ InjOn g (Iio (ord <| succ #α)) := by have h := lift_mk_le_lift_mk_of_injective <| injOn_iff_injective.1 h_inj have mk_initialSeg_subtype : #(Iio (ord <| succ #α)) = lift.{u + 1} (succ #α) := by - simpa only [coe_setOf, card_typein, card_ord] using mk_initialSeg (ord <| succ #α) + simpa only [coe_setOf, card_typein, card_ord] using mk_Iio_ordinal (ord <| succ #α) rw [mk_initialSeg_subtype, lift_lift, lift_le] at h exact not_le_of_lt (Order.lt_succ #α) h From 44c18bb2e11f833fcb719529f96479c40800d6cd Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Mon, 7 Oct 2024 17:59:13 +0000 Subject: [PATCH 28/43] chore(SetTheory/Cardinal/Ordinal): deprecate `ord_aleph_eq_enum_card` and friends (#17077) See [Zulip](https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/Deprecate.20.60ord_aleph'_eq_enum_card.60/near/472441290). --- Mathlib/SetTheory/Cardinal/Ordinal.lean | 17 +++++++++++++++++ 1 file changed, 17 insertions(+) diff --git a/Mathlib/SetTheory/Cardinal/Ordinal.lean b/Mathlib/SetTheory/Cardinal/Ordinal.lean index fe8a2ac903837..53142d747495b 100644 --- a/Mathlib/SetTheory/Cardinal/Ordinal.lean +++ b/Mathlib/SetTheory/Cardinal/Ordinal.lean @@ -325,7 +325,17 @@ theorem aleph0_lt_aleph_one : ℵ₀ < aleph 1 := by theorem countable_iff_lt_aleph_one {α : Type*} (s : Set α) : s.Countable ↔ #s < aleph 1 := by rw [← succ_aleph0, lt_succ_iff, le_aleph0_iff_set_countable] +section deprecated + +set_option linter.deprecated false + +-- TODO: these lemmas should be stated in terms of the `ω` function and of an `IsInitial` predicate, +-- neither of which currently exist. +-- +-- They should also use `¬ BddAbove` instead of `Unbounded (· < ·)`. + /-- Ordinals that are cardinals are unbounded. -/ +@[deprecated (since := "2024-09-24")] theorem ord_card_unbounded : Unbounded (· < ·) { b : Ordinal | b.card.ord = b } := unbounded_lt_iff.2 fun a => ⟨_, @@ -333,10 +343,12 @@ theorem ord_card_unbounded : Unbounded (· < ·) { b : Ordinal | b.card.ord = b dsimp rw [card_ord], (lt_ord_succ_card a).le⟩⟩ +@[deprecated (since := "2024-09-24")] theorem eq_aleph'_of_eq_card_ord {o : Ordinal} (ho : o.card.ord = o) : ∃ a, (aleph' a).ord = o := ⟨aleph'.symm o.card, by simpa using ho⟩ /-- `ord ∘ aleph'` enumerates the ordinals that are cardinals. -/ +@[deprecated (since := "2024-09-24")] theorem ord_aleph'_eq_enum_card : ord ∘ aleph' = enumOrd { b : Ordinal | b.card.ord = b } := by rw [← eq_enumOrd _ ord_card_unbounded, range_eq_iff] exact @@ -346,9 +358,11 @@ theorem ord_aleph'_eq_enum_card : ord ∘ aleph' = enumOrd { b : Ordinal | b.car rw [card_ord], fun b hb => eq_aleph'_of_eq_card_ord hb⟩⟩ /-- Infinite ordinals that are cardinals are unbounded. -/ +@[deprecated (since := "2024-09-24")] theorem ord_card_unbounded' : Unbounded (· < ·) { b : Ordinal | b.card.ord = b ∧ ω ≤ b } := (unbounded_lt_inter_le ω).2 ord_card_unbounded +@[deprecated (since := "2024-09-24")] theorem eq_aleph_of_eq_card_ord {o : Ordinal} (ho : o.card.ord = o) (ho' : ω ≤ o) : ∃ a, (aleph a).ord = o := by cases' eq_aleph'_of_eq_card_ord ho with a ha @@ -357,6 +371,7 @@ theorem eq_aleph_of_eq_card_ord {o : Ordinal} (ho : o.card.ord = o) (ho' : ω rwa [← aleph0_le_aleph', ← ord_le_ord, ha, ord_aleph0] /-- `ord ∘ aleph` enumerates the infinite ordinals that are cardinals. -/ +@[deprecated (since := "2024-09-24")] theorem ord_aleph_eq_enum_card : ord ∘ aleph = enumOrd { b : Ordinal | b.card.ord = b ∧ ω ≤ b } := by rw [← eq_enumOrd _ ord_card_unbounded'] @@ -367,6 +382,8 @@ theorem ord_aleph_eq_enum_card : · rw [← ord_aleph0, Function.comp_apply, ord_le_ord] exact aleph0_le_aleph _ +end deprecated + end aleph /-! ### Beth cardinals -/ From 71c35205aaec91dfdf04da8b087ff6f1f4a2053f Mon Sep 17 00:00:00 2001 From: Bhavik Mehta Date: Mon, 7 Oct 2024 17:59:14 +0000 Subject: [PATCH 29/43] feat(Data/Set): monotonicity on insert (#17086) --- Mathlib/Data/Set/Monotone.lean | 20 ++++++++++++++++++++ 1 file changed, 20 insertions(+) diff --git a/Mathlib/Data/Set/Monotone.lean b/Mathlib/Data/Set/Monotone.lean index f811b7ae15131..5bd60bb9fb1c5 100644 --- a/Mathlib/Data/Set/Monotone.lean +++ b/Mathlib/Data/Set/Monotone.lean @@ -85,6 +85,16 @@ protected theorem _root_.StrictAntiOn.strictAnti (h : StrictAntiOn f s) : StrictAnti (f ∘ Subtype.val : s → β) := fun x y hlt => h x.coe_prop y.coe_prop hlt +lemma MonotoneOn_insert_iff {a : α} : + MonotoneOn f (insert a s) ↔ + (∀ b ∈ s, b ≤ a → f b ≤ f a) ∧ (∀ b ∈ s, a ≤ b → f a ≤ f b) ∧ MonotoneOn f s := by + simp [MonotoneOn, forall_and] + +lemma AntitoneOn_insert_iff {a : α} : + AntitoneOn f (insert a s) ↔ + (∀ b ∈ s, b ≤ a → f a ≤ f b) ∧ (∀ b ∈ s, a ≤ b → f b ≤ f a) ∧ AntitoneOn f s := + @MonotoneOn_insert_iff α βᵒᵈ _ _ _ _ _ + end Mono end Set @@ -148,6 +158,16 @@ theorem StrictMono.codRestrict [Preorder α] [Preorder β] {f : α → β} (hf : {s : Set β} (hs : ∀ x, f x ∈ s) : StrictMono (Set.codRestrict f s hs) := hf +lemma strictMonoOn_insert_iff [Preorder α] [Preorder β] {f : α → β} {s : Set α} {a : α} : + StrictMonoOn f (insert a s) ↔ + (∀ b ∈ s, b < a → f b < f a) ∧ (∀ b ∈ s, a < b → f a < f b) ∧ StrictMonoOn f s := by + simp [StrictMonoOn, forall_and] + +lemma strictAntiOn_insert_iff [Preorder α] [Preorder β] {f : α → β} {s : Set α} {a : α} : + StrictAntiOn f (insert a s) ↔ + (∀ b ∈ s, b < a → f a < f b) ∧ (∀ b ∈ s, a < b → f b < f a) ∧ StrictAntiOn f s := + @strictMonoOn_insert_iff α βᵒᵈ _ _ _ _ _ + end strictMono namespace Function From 6c2b2f49b08f96e1d598b5143e3f86d011ee07f1 Mon Sep 17 00:00:00 2001 From: damiano Date: Mon, 7 Oct 2024 17:59:15 +0000 Subject: [PATCH 30/43] chore: some more cdots (#17402) --- .../Analysis/Calculus/FDeriv/Analytic.lean | 6 ++-- Mathlib/Data/Int/WithZero.lean | 4 ++- .../RingedSpace/PresheafedSpace/Gluing.lean | 34 +++++++++---------- Mathlib/GroupTheory/GroupAction/Quotient.lean | 2 +- 4 files changed, 24 insertions(+), 22 deletions(-) diff --git a/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean b/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean index 0e90e3c0a221f..80735fc0538cf 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean @@ -361,9 +361,9 @@ protected theorem AnalyticOn.fderivWithin (h : AnalyticOn 𝕜 f s) (hu : Unique refine ⟨p.derivSeries, r, ?_⟩ refine ⟨hr.r_le.trans p.radius_le_radius_derivSeries, hr.r_pos, fun {y} hy h'y ↦ ?_⟩ apply hr.hasSum_derivSeries_of_hasFDerivWithinAt (by simpa [edist_eq_coe_nnnorm] using h'y) hy - rw [insert_eq_of_mem hx] at hy ⊢ - apply DifferentiableWithinAt.hasFDerivWithinAt - · exact h.differentiableOn _ hy + · rw [insert_eq_of_mem hx] at hy ⊢ + apply DifferentiableWithinAt.hasFDerivWithinAt + exact h.differentiableOn _ hy · rwa [insert_eq_of_mem hx] /-- If a function is analytic on a set `s`, so are its successive Fréchet derivative within this diff --git a/Mathlib/Data/Int/WithZero.lean b/Mathlib/Data/Int/WithZero.lean index 4e3294a68595e..3cf41e5f88e38 100644 --- a/Mathlib/Data/Int/WithZero.lean +++ b/Mathlib/Data/Int/WithZero.lean @@ -63,7 +63,9 @@ theorem toNNReal_pos_apply {e : NNReal} (he : e ≠ 0) {x : ℤₘ₀} (hx : x = theorem toNNReal_neg_apply {e : NNReal} (he : e ≠ 0) {x : ℤₘ₀} (hx : x ≠ 0) : toNNReal he x = e ^ Multiplicative.toAdd (WithZero.unzero hx) := by simp only [toNNReal, MonoidWithZeroHom.coe_mk, ZeroHom.coe_mk] - split_ifs; tauto; rfl + split_ifs + · tauto + · rfl /-- `toNNReal` sends nonzero elements to nonzero elements. -/ theorem toNNReal_ne_zero {e : NNReal} {m : ℤₘ₀} (he : e ≠ 0) (hm : m ≠ 0) : toNNReal he m ≠ 0 := by diff --git a/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean b/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean index 4d9076090a048..b84225323dd1d 100644 --- a/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean +++ b/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean @@ -439,23 +439,23 @@ theorem π_ιInvApp_π (i j : D.J) (U : Opens (D.U i).carrier) : rw [← @cancel_mono (f := (componentwiseDiagram 𝖣.diagram.multispan _).map (Quiver.Hom.op (WalkingMultispan.Hom.snd (i, j))) ≫ 𝟙 _) ..] - simp_rw [Category.assoc] - rw [limit.w_assoc] - erw [limit.lift_π_assoc] - rw [Category.comp_id, Category.comp_id] - change _ ≫ _ ≫ (_ ≫ _) ≫ _ = _ - rw [congr_app (D.t_id _), id_c_app] - simp_rw [Category.assoc] - rw [← Functor.map_comp_assoc] - -- Porting note (#11224): change `rw` to `erw` - erw [IsOpenImmersion.inv_naturality_assoc] - erw [IsOpenImmersion.app_invApp_assoc] - iterate 3 rw [← Functor.map_comp_assoc] - rw [NatTrans.naturality_assoc] - erw [← (D.V (i, j)).presheaf.map_comp] - convert - limit.w (componentwiseDiagram 𝖣.diagram.multispan _) - (Quiver.Hom.op (WalkingMultispan.Hom.fst (i, j))) + · simp_rw [Category.assoc] + rw [limit.w_assoc] + erw [limit.lift_π_assoc] + rw [Category.comp_id, Category.comp_id] + change _ ≫ _ ≫ (_ ≫ _) ≫ _ = _ + rw [congr_app (D.t_id _), id_c_app] + simp_rw [Category.assoc] + rw [← Functor.map_comp_assoc] + -- Porting note (#11224): change `rw` to `erw` + erw [IsOpenImmersion.inv_naturality_assoc] + erw [IsOpenImmersion.app_invApp_assoc] + iterate 3 rw [← Functor.map_comp_assoc] + rw [NatTrans.naturality_assoc] + erw [← (D.V (i, j)).presheaf.map_comp] + convert + limit.w (componentwiseDiagram 𝖣.diagram.multispan _) + (Quiver.Hom.op (WalkingMultispan.Hom.fst (i, j))) · rw [Category.comp_id] apply (config := { allowSynthFailures := true }) mono_comp change Mono ((_ ≫ D.f j i).c.app _) diff --git a/Mathlib/GroupTheory/GroupAction/Quotient.lean b/Mathlib/GroupTheory/GroupAction/Quotient.lean index 7bf00ac53f9d3..2cdd8bb08b134 100644 --- a/Mathlib/GroupTheory/GroupAction/Quotient.lean +++ b/Mathlib/GroupTheory/GroupAction/Quotient.lean @@ -402,7 +402,7 @@ noncomputable def equivSubgroupOrbitsQuotientGroup [IsPretransitive α β] rw [Quotient.eq'', leftRel_eq] simp only convert one_mem H - rw [inv_mul_eq_one, eq_comm, ← inv_mul_eq_one, ← Subgroup.mem_bot, ← free (g⁻¹ • x), + · rw [inv_mul_eq_one, eq_comm, ← inv_mul_eq_one, ← Subgroup.mem_bot, ← free (g⁻¹ • x), mem_stabilizer_iff, mul_smul, (exists_smul_eq α (g⁻¹ • x) x).choose_spec] #adaptation_note /-- From 10671b29c9208b978c91d72721fbcca39b3c2c4a Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 7 Oct 2024 17:59:16 +0000 Subject: [PATCH 31/43] chore: generalise lemmas from `OrderedSemiring` to `MonoidWithZero` (#17412) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit and add the `₀` to follow the naming convention --- Archive/Imo/Imo1962Q1.lean | 2 +- Archive/Wiedijk100Theorems/AbelRuffini.lean | 4 +- .../SumOfPrimeReciprocalsDiverges.lean | 2 +- Mathlib/Algebra/Order/Field/Basic.lean | 2 +- .../Order/GroupWithZero/Canonical.lean | 5 -- .../Order/GroupWithZero/Unbundled.lean | 46 +++++++++++++++++-- .../Algebra/Order/Interval/Set/Instances.lean | 7 +-- Mathlib/Algebra/Order/Ring/Basic.lean | 45 +++--------------- .../Polynomial/Degree/CardPowDegree.lean | 2 +- .../FundamentalGroupoid/Basic.lean | 4 +- Mathlib/Analysis/Analytic/Composition.lean | 2 +- Mathlib/Analysis/Convex/Between.lean | 6 +-- .../Analysis/Distribution/SchwartzSpace.lean | 4 +- .../Fourier/FourierTransformDeriv.lean | 2 +- .../LocallyConvex/BalancedCoreHull.lean | 2 +- Mathlib/Analysis/Normed/Field/UnitBall.lean | 2 +- Mathlib/Analysis/PSeries.lean | 6 +-- .../SpecialFunctions/Pow/Continuity.lean | 2 +- Mathlib/Analysis/SpecificLimits/Basic.lean | 4 +- Mathlib/Analysis/SpecificLimits/FloorPow.lean | 4 +- Mathlib/Combinatorics/SetFamily/Kleitman.lean | 2 +- .../Combinatorics/SimpleGraph/Density.lean | 2 +- .../SimpleGraph/Regularity/Chunk.lean | 2 +- .../SimpleGraph/Regularity/Lemma.lean | 2 +- Mathlib/Computability/Ackermann.lean | 3 +- Mathlib/Data/Complex/Exponential.lean | 2 +- Mathlib/Data/ENNReal/Inv.lean | 4 +- Mathlib/Data/Nat/Choose/Factorization.lean | 2 +- Mathlib/Data/Real/Pi/Bounds.lean | 2 +- Mathlib/Data/Real/Pi/Irrational.lean | 2 +- Mathlib/NumberTheory/Bertrand.lean | 4 +- .../ClassNumber/AdmissibleCardPowDegree.lean | 2 +- Mathlib/NumberTheory/FermatPsp.lean | 2 +- .../Liouville/LiouvilleNumber.lean | 2 +- Mathlib/NumberTheory/Padics/Hensel.lean | 8 ++-- .../NumberTheory/Padics/PadicIntegers.lean | 2 +- Mathlib/Order/Filter/AtTopBot/Ring.lean | 4 +- Mathlib/Probability/StrongLaw.lean | 2 +- Mathlib/Topology/Algebra/PontryaginDual.lean | 2 +- Mathlib/Topology/UnitInterval.lean | 2 +- 40 files changed, 104 insertions(+), 102 deletions(-) diff --git a/Archive/Imo/Imo1962Q1.lean b/Archive/Imo/Imo1962Q1.lean index dd0ef63d96b8b..d151c6df95fc3 100644 --- a/Archive/Imo/Imo1962Q1.lean +++ b/Archive/Imo/Imo1962Q1.lean @@ -107,7 +107,7 @@ lemma case_more_digits {c n : ℕ} (hc : (digits 10 c).length ≥ 6) (hpp : Prob calc n ≥ 10 * c := le.intro hpp.left.symm _ ≥ 10 ^ (digits 10 c).length := base_pow_length_digits_le 10 c (by decide) hnz - _ ≥ 10 ^ 6 := pow_le_pow_right (by decide) hc + _ ≥ 10 ^ 6 := pow_right_mono₀ (by decide) hc _ ≥ 153846 := by norm_num /-! diff --git a/Archive/Wiedijk100Theorems/AbelRuffini.lean b/Archive/Wiedijk100Theorems/AbelRuffini.lean index df179eb859924..36026dc2ddd76 100644 --- a/Archive/Wiedijk100Theorems/AbelRuffini.lean +++ b/Archive/Wiedijk100Theorems/AbelRuffini.lean @@ -115,7 +115,7 @@ theorem real_roots_Phi_ge_aux (hab : b < a) : · have hf1 : f 1 < 0 := by simp [hf, hb] have hfa : 0 ≤ f a := by simp_rw [hf, ← sq] - refine add_nonneg (sub_nonneg.mpr (pow_le_pow_right ha ?_)) ?_ <;> norm_num + refine add_nonneg (sub_nonneg.mpr (pow_right_mono₀ ha ?_)) ?_ <;> norm_num obtain ⟨x, ⟨-, hx1⟩, hx2⟩ := intermediate_value_Ico' hle (hc _) (Set.mem_Ioc.mpr ⟨hf1, hf0⟩) obtain ⟨y, ⟨hy1, -⟩, hy2⟩ := intermediate_value_Ioc ha (hc _) (Set.mem_Ioc.mpr ⟨hf1, hfa⟩) exact ⟨x, y, (hx1.trans hy1).ne, hx2, hy2⟩ @@ -126,7 +126,7 @@ theorem real_roots_Phi_ge_aux (hab : b < a) : f (-a) = (a : ℝ) ^ 2 - (a : ℝ) ^ 5 + b := by norm_num [hf, ← sq, sub_eq_add_neg, add_comm, Odd.neg_pow (by decide : Odd 5)] _ ≤ (a : ℝ) ^ 2 - (a : ℝ) ^ 3 + (a - 1) := by - refine add_le_add (sub_le_sub_left (pow_le_pow_right ha ?_) _) ?_ <;> linarith + refine add_le_add (sub_le_sub_left (pow_right_mono₀ ha ?_) _) ?_ <;> linarith _ = -((a : ℝ) - 1) ^ 2 * (a + 1) := by ring _ ≤ 0 := by nlinarith have ha' := neg_nonpos.mpr (hle.trans ha) diff --git a/Archive/Wiedijk100Theorems/SumOfPrimeReciprocalsDiverges.lean b/Archive/Wiedijk100Theorems/SumOfPrimeReciprocalsDiverges.lean index 74017293ce14b..e30881c10059f 100644 --- a/Archive/Wiedijk100Theorems/SumOfPrimeReciprocalsDiverges.lean +++ b/Archive/Wiedijk100Theorems/SumOfPrimeReciprocalsDiverges.lean @@ -154,7 +154,7 @@ theorem card_le_two_pow {x k : ℕ} : card M₁ ≤ card (image f K) := card_le_card h _ ≤ card K := card_image_le _ ≤ 2 ^ card (image Nat.succ (range k)) := by simp only [K, card_powerset]; rfl - _ ≤ 2 ^ card (range k) := pow_le_pow_right one_le_two card_image_le + _ ≤ 2 ^ card (range k) := pow_right_mono₀ one_le_two card_image_le _ = 2 ^ k := by rw [card_range k] /-- diff --git a/Mathlib/Algebra/Order/Field/Basic.lean b/Mathlib/Algebra/Order/Field/Basic.lean index 7121bb47a6885..b349a1c749325 100644 --- a/Mathlib/Algebra/Order/Field/Basic.lean +++ b/Mathlib/Algebra/Order/Field/Basic.lean @@ -415,7 +415,7 @@ theorem one_div_strictAntiOn : StrictAntiOn (fun x : α => 1 / x) (Set.Ioi 0) := theorem one_div_pow_le_one_div_pow_of_le (a1 : 1 ≤ a) {m n : ℕ} (mn : m ≤ n) : 1 / a ^ n ≤ 1 / a ^ m := by - refine (one_div_le_one_div ?_ ?_).mpr (pow_le_pow_right a1 mn) <;> + refine (one_div_le_one_div ?_ ?_).mpr (pow_right_mono₀ a1 mn) <;> exact pow_pos (zero_lt_one.trans_le a1) _ theorem one_div_pow_lt_one_div_pow_of_lt (a1 : 1 < a) {m n : ℕ} (mn : m < n) : diff --git a/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean b/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean index 0180fe8cbd33f..5809731874601 100644 --- a/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean +++ b/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean @@ -130,11 +130,6 @@ instance (priority := 100) LinearOrderedCommGroupWithZero.toMulPosStrictMono : MulPosStrictMono α where elim a b c hbc := by by_contra! h; exact hbc.not_le <| (mul_le_mul_right a.2).1 h -/-- Alias of `mul_le_one'` for unification. -/ -@[deprecated mul_le_one' (since := "2024-08-21")] -theorem mul_le_one₀ (ha : a ≤ 1) (hb : b ≤ 1) : a * b ≤ 1 := - mul_le_one' ha hb - /-- Alias of `one_le_mul'` for unification. -/ @[deprecated one_le_mul (since := "2024-08-21")] theorem one_le_mul₀ (ha : 1 ≤ a) (hb : 1 ≤ b) : 1 ≤ a * b := diff --git a/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean b/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean index a83a71f9e5b2e..54e2ea4cb2615 100644 --- a/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean +++ b/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean @@ -935,12 +935,16 @@ section MonoidWithZero variable [MonoidWithZero M₀] section Preorder -variable [Preorder M₀] {a b c d : M₀} {n : ℕ} +variable [Preorder M₀] {a b c d : M₀} {m n : ℕ} @[simp] lemma pow_nonneg [ZeroLEOneClass M₀] [PosMulMono M₀] (ha : 0 ≤ a) : ∀ n, 0 ≤ a ^ n | 0 => pow_zero a ▸ zero_le_one | n + 1 => pow_succ a n ▸ mul_nonneg (pow_nonneg ha _) ha +lemma zero_pow_le_one [ZeroLEOneClass M₀] : ∀ n : ℕ, (0 : M₀) ^ n ≤ 1 + | 0 => (pow_zero _).le + | n + 1 => by rw [zero_pow n.succ_ne_zero]; exact zero_le_one + lemma pow_le_pow_of_le_one [ZeroLEOneClass M₀] [PosMulMono M₀] [MulPosMono M₀] (ha₀ : 0 ≤ a) (ha₁ : a ≤ 1) : ∀ {m n : ℕ}, m ≤ n → a ^ n ≤ a ^ m | _, _, Nat.le.refl => le_rfl @@ -958,9 +962,6 @@ lemma sq_le [ZeroLEOneClass M₀] [PosMulMono M₀] [MulPosMono M₀] (h₀ : 0 lemma one_le_mul_of_one_le_of_one_le [ZeroLEOneClass M₀] [PosMulMono M₀] (ha : 1 ≤ a) (hb : 1 ≤ b) : (1 : M₀) ≤ a * b := Left.one_le_mul_of_le_of_le ha hb <| zero_le_one.trans ha -lemma mul_le_one [ZeroLEOneClass M₀] [PosMulMono M₀] [MulPosMono M₀] (ha : a ≤ 1) (hb₀ : 0 ≤ b) - (hb : b ≤ 1) : a * b ≤ 1 := one_mul (1 : M₀) ▸ mul_le_mul ha hb hb₀ zero_le_one - lemma one_lt_mul_of_le_of_lt [ZeroLEOneClass M₀] [MulPosMono M₀] (ha : 1 ≤ a) (hb : 1 < b) : 1 < a * b := hb.trans_le <| le_mul_of_one_le_left (zero_le_one.trans hb.le) ha @@ -975,6 +976,43 @@ lemma mul_lt_one_of_nonneg_of_lt_one_left [PosMulMono M₀] (ha₀ : 0 ≤ a) (h lemma mul_lt_one_of_nonneg_of_lt_one_right [MulPosMono M₀] (ha : a ≤ 1) (hb₀ : 0 ≤ b) (hb : b < 1) : a * b < 1 := (mul_le_of_le_one_left hb₀ ha).trans_lt hb +section +variable [ZeroLEOneClass M₀] [PosMulMono M₀] [MulPosMono M₀] + +lemma mul_le_one₀ (ha : a ≤ 1) (hb₀ : 0 ≤ b) (hb : b ≤ 1) : a * b ≤ 1 := + one_mul (1 : M₀) ▸ mul_le_mul ha hb hb₀ zero_le_one + +lemma pow_le_one₀ : ∀ {n : ℕ}, 0 ≤ a → a ≤ 1 → a ^ n ≤ 1 + | 0, _, _ => (pow_zero a).le + | n + 1, h₀, h₁ => (pow_succ a n).le.trans (mul_le_one₀ (pow_le_one₀ h₀ h₁) h₀ h₁) + +lemma pow_lt_one₀ (h₀ : 0 ≤ a) (h₁ : a < 1) : ∀ {n : ℕ}, n ≠ 0 → a ^ n < 1 + | 0, h => (h rfl).elim + | n + 1, _ => by + rw [pow_succ']; exact mul_lt_one_of_nonneg_of_lt_one_left h₀ h₁ (pow_le_one₀ h₀ h₁.le) + +lemma one_le_pow₀ (ha : 1 ≤ a) : ∀ {n : ℕ}, 1 ≤ a ^ n + | 0 => by rw [pow_zero] + | n + 1 => by + simpa only [pow_succ', mul_one] + using mul_le_mul ha (one_le_pow₀ ha) zero_le_one (zero_le_one.trans ha) + +lemma one_lt_pow₀ (ha : 1 < a) : ∀ {n : ℕ}, n ≠ 0 → 1 < a ^ n + | 0, h => (h rfl).elim + | n + 1, _ => by rw [pow_succ']; exact one_lt_mul_of_lt_of_le ha (one_le_pow₀ ha.le) + +lemma pow_right_mono₀ (h : 1 ≤ a) : Monotone (a ^ ·) := + monotone_nat_of_le_succ fun n => by + rw [pow_succ']; exact le_mul_of_one_le_left (pow_nonneg (zero_le_one.trans h) _) h + +@[gcongr] +lemma pow_le_pow_right₀ (ha : 1 ≤ a) (hmn : m ≤ n) : a ^ m ≤ a ^ n := pow_right_mono₀ ha hmn + +lemma le_self_pow₀ (ha : 1 ≤ a) (hn : n ≠ 0) : a ≤ a ^ n := by + simpa only [pow_one] using pow_le_pow_right₀ ha <| Nat.pos_iff_ne_zero.2 hn + +end + variable [Preorder α] {f g : α → M₀} lemma monotone_mul_left_of_nonneg [PosMulMono M₀] (ha : 0 ≤ a) : Monotone fun x ↦ a * x := diff --git a/Mathlib/Algebra/Order/Interval/Set/Instances.lean b/Mathlib/Algebra/Order/Interval/Set/Instances.lean index da5a81838c6e1..4624971afef4b 100644 --- a/Mathlib/Algebra/Order/Interval/Set/Instances.lean +++ b/Mathlib/Algebra/Order/Interval/Set/Instances.lean @@ -3,7 +3,8 @@ Copyright (c) 2022 Stuart Presnell. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Stuart Presnell, Eric Wieser, Yaël Dillies, Patrick Massot, Kim Morrison -/ -import Mathlib.Algebra.Order.Ring.Basic +import Mathlib.Algebra.GroupWithZero.InjSurj +import Mathlib.Algebra.Order.Ring.Defs import Mathlib.Algebra.Ring.Regular import Mathlib.Order.Interval.Set.Basic @@ -98,7 +99,7 @@ theorem le_one {t : Icc (0 : α) 1} : t ≤ 1 := t.2.2 instance mul : Mul (Icc (0 : α) 1) where - mul p q := ⟨p * q, ⟨mul_nonneg p.2.1 q.2.1, mul_le_one p.2.2 q.2.1 q.2.2⟩⟩ + mul p q := ⟨p * q, ⟨mul_nonneg p.2.1 q.2.1, mul_le_one₀ p.2.2 q.2.1 q.2.2⟩⟩ instance pow : Pow (Icc (0 : α) 1) ℕ where pow p n := ⟨p.1 ^ n, ⟨pow_nonneg p.2.1 n, pow_le_one₀ p.2.1 p.2.2⟩⟩ @@ -236,7 +237,7 @@ theorem le_one {t : Ioc (0 : α) 1} : t ≤ 1 := t.2.2 instance mul : Mul (Ioc (0 : α) 1) where - mul p q := ⟨p.1 * q.1, ⟨mul_pos p.2.1 q.2.1, mul_le_one p.2.2 (le_of_lt q.2.1) q.2.2⟩⟩ + mul p q := ⟨p.1 * q.1, ⟨mul_pos p.2.1 q.2.1, mul_le_one₀ p.2.2 (le_of_lt q.2.1) q.2.2⟩⟩ instance pow : Pow (Ioc (0 : α) 1) ℕ where pow p n := ⟨p.1 ^ n, ⟨pow_pos p.2.1 n, pow_le_one₀ (le_of_lt p.2.1) p.2.2⟩⟩ diff --git a/Mathlib/Algebra/Order/Ring/Basic.lean b/Mathlib/Algebra/Order/Ring/Basic.lean index 5a0821b136c4a..5e64f73cc40ea 100644 --- a/Mathlib/Algebra/Order/Ring/Basic.lean +++ b/Mathlib/Algebra/Order/Ring/Basic.lean @@ -37,10 +37,6 @@ section OrderedSemiring variable [OrderedSemiring R] {a b x y : R} {n m : ℕ} -theorem zero_pow_le_one : ∀ n : ℕ, (0 : R) ^ n ≤ 1 - | 0 => (pow_zero _).le - | n + 1 => by rw [zero_pow n.succ_ne_zero]; exact zero_le_one - theorem pow_add_pow_le (hx : 0 ≤ x) (hy : 0 ≤ y) (hn : n ≠ 0) : x ^ n + y ^ n ≤ (x + y) ^ n := by rcases Nat.exists_eq_add_one_of_ne_zero hn with ⟨k, rfl⟩ induction k with @@ -60,48 +56,21 @@ theorem pow_add_pow_le (hx : 0 ≤ x) (hy : 0 ≤ y) (hn : n ≠ 0) : x ^ n + y rw [pow_succ' _ n] exact mul_le_mul_of_nonneg_left (ih (Nat.succ_ne_zero k)) h2 -@[bound] -theorem pow_le_one₀ : ∀ {n : ℕ}, 0 ≤ a → a ≤ 1 → a ^ n ≤ 1 - | 0, _, _ => (pow_zero a).le - | n + 1, h₀, h₁ => (pow_succ a n).le.trans (mul_le_one (pow_le_one₀ h₀ h₁) h₀ h₁) - -theorem pow_lt_one₀ (h₀ : 0 ≤ a) (h₁ : a < 1) : ∀ {n : ℕ}, n ≠ 0 → a ^ n < 1 - | 0, h => (h rfl).elim - | n + 1, _ => by - rw [pow_succ'] - exact mul_lt_one_of_nonneg_of_lt_one_left h₀ h₁ (pow_le_one₀ h₀ h₁.le) - -@[bound] -theorem one_le_pow₀ (H : 1 ≤ a) : ∀ {n : ℕ}, 1 ≤ a ^ n - | 0 => by rw [pow_zero] - | n + 1 => by - simpa only [pow_succ', mul_one] - using mul_le_mul H (one_le_pow₀ H) zero_le_one (zero_le_one.trans H) - -lemma one_lt_pow₀ (ha : 1 < a) : ∀ {n : ℕ}, n ≠ 0 → 1 < a ^ n - | 0, h => (h rfl).elim - | n + 1, _ => by rw [pow_succ']; exact one_lt_mul_of_lt_of_le ha (one_le_pow₀ ha.le) +attribute [bound] pow_le_one₀ one_le_pow₀ +@[deprecated (since := "2024-09-28")] alias mul_le_one := mul_le_one₀ @[deprecated (since := "2024-09-28")] alias pow_le_one := pow_le_one₀ @[deprecated (since := "2024-09-28")] alias pow_lt_one := pow_lt_one₀ @[deprecated (since := "2024-09-28")] alias one_le_pow_of_one_le := one_le_pow₀ @[deprecated (since := "2024-09-28")] alias one_lt_pow := one_lt_pow₀ - -theorem pow_right_mono (h : 1 ≤ a) : Monotone (a ^ ·) := - monotone_nat_of_le_succ fun n => by - rw [pow_succ'] - exact le_mul_of_one_le_left (pow_nonneg (zero_le_one.trans h) _) h - -@[gcongr] -theorem pow_le_pow_right (ha : 1 ≤ a) (h : n ≤ m) : a ^ n ≤ a ^ m := pow_right_mono ha h - -theorem le_self_pow (ha : 1 ≤ a) (h : m ≠ 0) : a ≤ a ^ m := by - simpa only [pow_one] using pow_le_pow_right ha <| Nat.pos_iff_ne_zero.2 h +@[deprecated (since := "2024-10-04")] alias pow_right_mono := pow_right_mono₀ +@[deprecated (since := "2024-10-04")] alias pow_le_pow_right := pow_le_pow_right₀ +@[deprecated (since := "2024-10-04")] alias le_self_pow := le_self_pow₀ /-- The `bound` tactic can't handle `m ≠ 0` goals yet, so we express as `0 < m` -/ @[bound] lemma Bound.le_self_pow_of_pos {m : ℕ} (ha : 1 ≤ a) (h : 0 < m) : a ≤ a ^ m := - le_self_pow ha h.ne' + le_self_pow₀ ha h.ne' @[mono, gcongr, bound] theorem pow_le_pow_left {a b : R} (ha : 0 ≤ a) (hab : a ≤ b) : ∀ n, a ^ n ≤ b ^ n @@ -119,7 +88,7 @@ lemma pow_add_pow_le' (ha : 0 ≤ a) (hb : 0 ≤ b) : a ^ n + b ^ n ≤ 2 * (a + lemma Bound.pow_le_pow_right_of_le_one_or_one_le (h : 1 ≤ a ∧ n ≤ m ∨ 0 ≤ a ∧ a ≤ 1 ∧ m ≤ n) : a ^ n ≤ a ^ m := by rcases h with ⟨a1, nm⟩ | ⟨a0, a1, mn⟩ - · exact pow_le_pow_right a1 nm + · exact pow_right_mono₀ a1 nm · exact pow_le_pow_of_le_one a0 a1 mn end OrderedSemiring diff --git a/Mathlib/Algebra/Polynomial/Degree/CardPowDegree.lean b/Mathlib/Algebra/Polynomial/Degree/CardPowDegree.lean index 01b1c8ec0561e..8b5296df327a3 100644 --- a/Mathlib/Algebra/Polynomial/Degree/CardPowDegree.lean +++ b/Mathlib/Algebra/Polynomial/Degree/CardPowDegree.lean @@ -61,7 +61,7 @@ noncomputable def cardPowDegree : AbsoluteValue Fq[X] ℤ := · simp only [hpq, hp, hq, eq_self_iff_true, if_true, if_false] exact add_nonneg (pow_pos _).le (pow_pos _).le simp only [hpq, hp, hq, if_false] - refine le_trans (pow_le_pow_right (by omega) (Polynomial.natDegree_add_le _ _)) ?_ + refine le_trans (pow_right_mono₀ (by omega) (Polynomial.natDegree_add_le _ _)) ?_ refine le_trans (le_max_iff.mpr ?_) (max_le_add_of_nonneg (pow_nonneg (by omega) _) (pow_nonneg (by omega) _)) diff --git a/Mathlib/AlgebraicTopology/FundamentalGroupoid/Basic.lean b/Mathlib/AlgebraicTopology/FundamentalGroupoid/Basic.lean index 84980eaca3178..2a8bb0c418f7e 100644 --- a/Mathlib/AlgebraicTopology/FundamentalGroupoid/Basic.lean +++ b/Mathlib/AlgebraicTopology/FundamentalGroupoid/Basic.lean @@ -59,7 +59,7 @@ theorem reflTransSymmAux_mem_I (x : I × I) : reflTransSymmAux x ∈ I := by · norm_num · unit_interval · rw [mul_assoc] - apply mul_le_one + apply mul_le_one₀ · unit_interval · apply mul_nonneg · norm_num @@ -69,7 +69,7 @@ theorem reflTransSymmAux_mem_I (x : I × I) : reflTransSymmAux x ∈ I := by · apply mul_nonneg · unit_interval linarith [unitInterval.nonneg x.2, unitInterval.le_one x.2] - · apply mul_le_one + · apply mul_le_one₀ · unit_interval · linarith [unitInterval.nonneg x.2, unitInterval.le_one x.2] · linarith [unitInterval.nonneg x.2, unitInterval.le_one x.2] diff --git a/Mathlib/Analysis/Analytic/Composition.lean b/Mathlib/Analysis/Analytic/Composition.lean index 49ef36b0ddbd9..c73176c9a5ec5 100644 --- a/Mathlib/Analysis/Analytic/Composition.lean +++ b/Mathlib/Analysis/Analytic/Composition.lean @@ -473,7 +473,7 @@ theorem comp_summable_nnreal (q : FormalMultilinearSeries 𝕜 F G) (p : FormalM simp only [Finset.prod_mul_distrib, Finset.prod_pow_eq_pow_sum, c.sum_blocksFun] _ ≤ ∏ _i : Fin c.length, Cp := Finset.prod_le_prod' fun i _ => hCp _ _ = Cp ^ c.length := by simp - _ ≤ Cp ^ n := pow_le_pow_right hCp1 c.length_le + _ ≤ Cp ^ n := pow_right_mono₀ hCp1 c.length_le calc ‖q.compAlongComposition p c‖₊ * r ^ n ≤ (‖q c.length‖₊ * ∏ i, ‖p (c.blocksFun i)‖₊) * r ^ n := diff --git a/Mathlib/Analysis/Convex/Between.lean b/Mathlib/Analysis/Convex/Between.lean index f9d6c966bcb5f..69429999e97e5 100644 --- a/Mathlib/Analysis/Convex/Between.lean +++ b/Mathlib/Analysis/Convex/Between.lean @@ -392,7 +392,7 @@ theorem sbtw_one_zero_iff {x : R} : Sbtw R 1 x 0 ↔ x ∈ Set.Ioo (0 : R) 1 := theorem Wbtw.trans_left {w x y z : P} (h₁ : Wbtw R w y z) (h₂ : Wbtw R w x y) : Wbtw R w x z := by rcases h₁ with ⟨t₁, ht₁, rfl⟩ rcases h₂ with ⟨t₂, ht₂, rfl⟩ - refine ⟨t₂ * t₁, ⟨mul_nonneg ht₂.1 ht₁.1, mul_le_one ht₂.2 ht₁.1 ht₁.2⟩, ?_⟩ + refine ⟨t₂ * t₁, ⟨mul_nonneg ht₂.1 ht₁.1, mul_le_one₀ ht₂.2 ht₁.1 ht₁.2⟩, ?_⟩ rw [lineMap_apply, lineMap_apply, lineMap_vsub_left, smul_smul] theorem Wbtw.trans_right {w x y z : P} (h₁ : Wbtw R w x z) (h₂ : Wbtw R x y z) : Wbtw R w y z := by @@ -695,8 +695,8 @@ theorem Wbtw.trans_left_right {w x y z : P} (h₁ : Wbtw R w y z) (h₂ : Wbtw R refine ⟨(t₁ - t₂ * t₁) / (1 - t₂ * t₁), ⟨div_nonneg (sub_nonneg.2 (mul_le_of_le_one_left ht₁.1 ht₂.2)) - (sub_nonneg.2 (mul_le_one ht₂.2 ht₁.1 ht₁.2)), - div_le_one_of_le₀ (sub_le_sub_right ht₁.2 _) (sub_nonneg.2 (mul_le_one ht₂.2 ht₁.1 ht₁.2))⟩, + (sub_nonneg.2 (mul_le_one₀ ht₂.2 ht₁.1 ht₁.2)), div_le_one_of_le₀ + (sub_le_sub_right ht₁.2 _) (sub_nonneg.2 (mul_le_one₀ ht₂.2 ht₁.1 ht₁.2))⟩, ?_⟩ simp only [lineMap_apply, smul_smul, ← add_vadd, vsub_vadd_eq_vsub_sub, smul_sub, ← sub_smul, ← add_smul, vadd_vsub, vadd_right_cancel_iff, div_mul_eq_mul_div, div_sub_div_same] diff --git a/Mathlib/Analysis/Distribution/SchwartzSpace.lean b/Mathlib/Analysis/Distribution/SchwartzSpace.lean index e7a99ac85a5d0..4e015c579727c 100644 --- a/Mathlib/Analysis/Distribution/SchwartzSpace.lean +++ b/Mathlib/Analysis/Distribution/SchwartzSpace.lean @@ -870,8 +870,8 @@ def compCLM {g : D → E} (hg : g.HasTemperateGrowth) rw [mul_pow] have hN₁' := (lt_of_lt_of_le zero_lt_one hN₁).ne' gcongr - · exact le_trans (by simp [hC]) (le_self_pow (by simp [hC]) hN₁') - · refine le_self_pow (one_le_pow₀ ?_) hN₁' + · exact le_trans (by simp [hC]) (le_self_pow₀ (by simp [hC]) hN₁') + · refine le_self_pow₀ (one_le_pow₀ ?_) hN₁' simp only [le_add_iff_nonneg_right, norm_nonneg] have := norm_iteratedFDeriv_comp_le f.smooth' hg.1 le_top x hbound hgrowth' have hxk : ‖x‖ ^ k ≤ (1 + ‖x‖) ^ k := diff --git a/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean b/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean index 16957cece15d7..0fbe303edbe46 100644 --- a/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean +++ b/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean @@ -413,7 +413,7 @@ lemma norm_iteratedFDeriv_fourierPowSMulRight · rw [← Nat.cast_pow, Nat.cast_le] calc n.descFactorial i ≤ n ^ i := Nat.descFactorial_le_pow _ _ _ ≤ (n + 1) ^ i := pow_le_pow_left (by omega) (by omega) i - _ ≤ (n + 1) ^ k := pow_le_pow_right (by omega) (Finset.mem_range_succ_iff.mp hi) + _ ≤ (n + 1) ^ k := pow_right_mono₀ (by omega) (Finset.mem_range_succ_iff.mp hi) · exact hv _ (by omega) _ (by omega) _ = (2 * n + 2) ^ k * (‖L‖^n * C) := by simp only [← Finset.sum_mul, ← Nat.cast_sum, Nat.sum_range_choose, mul_one, ← mul_assoc, diff --git a/Mathlib/Analysis/LocallyConvex/BalancedCoreHull.lean b/Mathlib/Analysis/LocallyConvex/BalancedCoreHull.lean index 5667478982a21..850c1e5f491d0 100644 --- a/Mathlib/Analysis/LocallyConvex/BalancedCoreHull.lean +++ b/Mathlib/Analysis/LocallyConvex/BalancedCoreHull.lean @@ -131,7 +131,7 @@ theorem balancedHull.balanced (s : Set E) : Balanced 𝕜 (balancedHull 𝕜 s) simp_rw [balancedHull, smul_set_iUnion₂, subset_def, mem_iUnion₂] rintro x ⟨r, hr, hx⟩ rw [← smul_assoc] at hx - exact ⟨a • r, (SeminormedRing.norm_mul _ _).trans (mul_le_one ha (norm_nonneg r) hr), hx⟩ + exact ⟨a • r, (SeminormedRing.norm_mul _ _).trans (mul_le_one₀ ha (norm_nonneg r) hr), hx⟩ end Module diff --git a/Mathlib/Analysis/Normed/Field/UnitBall.lean b/Mathlib/Analysis/Normed/Field/UnitBall.lean index 2817b533b7761..146e37b0e5af2 100644 --- a/Mathlib/Analysis/Normed/Field/UnitBall.lean +++ b/Mathlib/Analysis/Normed/Field/UnitBall.lean @@ -51,7 +51,7 @@ def Subsemigroup.unitClosedBall (𝕜 : Type*) [NonUnitalSeminormedRing 𝕜] : carrier := closedBall 0 1 mul_mem' hx hy := by rw [mem_closedBall_zero_iff] at * - exact (norm_mul_le _ _).trans (mul_le_one hx (norm_nonneg _) hy) + exact (norm_mul_le _ _).trans (mul_le_one₀ hx (norm_nonneg _) hy) instance Metric.unitClosedBall.semigroup [NonUnitalSeminormedRing 𝕜] : Semigroup (closedBall (0 : 𝕜) 1) := diff --git a/Mathlib/Analysis/PSeries.lean b/Mathlib/Analysis/PSeries.lean index aed28d9ca5915..1f09446ae9b7c 100644 --- a/Mathlib/Analysis/PSeries.lean +++ b/Mathlib/Analysis/PSeries.lean @@ -63,7 +63,7 @@ theorem le_sum_schlomilch' (hf : ∀ ⦃m n⦄, 0 < m → m ≤ n → f n ≤ f theorem le_sum_condensed' (hf : ∀ ⦃m n⦄, 0 < m → m ≤ n → f n ≤ f m) (n : ℕ) : (∑ k ∈ Ico 1 (2 ^ n), f k) ≤ ∑ k ∈ range n, 2 ^ k • f (2 ^ k) := by convert le_sum_schlomilch' hf (fun n => pow_pos zero_lt_two n) - (fun m n hm => pow_le_pow_right one_le_two hm) n using 2 + (fun m n hm => pow_right_mono₀ one_le_two hm) n using 2 simp [pow_succ, mul_two, two_mul] theorem le_sum_schlomilch (hf : ∀ ⦃m n⦄, 0 < m → m ≤ n → f n ≤ f m) (h_pos : ∀ n, 0 < u n) @@ -98,7 +98,7 @@ theorem sum_schlomilch_le' (hf : ∀ ⦃m n⦄, 1 < m → m ≤ n → f n ≤ f theorem sum_condensed_le' (hf : ∀ ⦃m n⦄, 1 < m → m ≤ n → f n ≤ f m) (n : ℕ) : (∑ k ∈ range n, 2 ^ k • f (2 ^ (k + 1))) ≤ ∑ k ∈ Ico 2 (2 ^ n + 1), f k := by convert sum_schlomilch_le' hf (fun n => pow_pos zero_lt_two n) - (fun m n hm => pow_le_pow_right one_le_two hm) n using 2 + (fun m n hm => pow_right_mono₀ one_le_two hm) n using 2 simp [pow_succ, mul_two, two_mul] theorem sum_schlomilch_le {C : ℕ} (hf : ∀ ⦃m n⦄, 1 < m → m ≤ n → f n ≤ f m) (h_pos : ∀ n, 0 < u n) @@ -416,7 +416,7 @@ theorem sum_Ioo_inv_sq_le (k n : ℕ) : (∑ i ∈ Ioo k n, (i ^ 2 : α)⁻¹) have A : (1 : α) ≤ k + 1 := by simp only [le_add_iff_nonneg_left, Nat.cast_nonneg] simp_rw [← one_div] gcongr - simpa using pow_le_pow_right A one_le_two + simpa using pow_right_mono₀ A one_le_two _ = 2 / (k + 1) := by ring end diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Continuity.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Continuity.lean index 529dea6298d19..d3ff288d66b92 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Continuity.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Continuity.lean @@ -394,7 +394,7 @@ theorem eventually_pow_one_div_le (x : ℝ≥0) {y : ℝ≥0} (hy : 1 < y) : refine eventually_atTop.2 ⟨m + 1, fun n hn => ?_⟩ simp only [one_div] simpa only [NNReal.rpow_inv_le_iff (Nat.cast_pos.2 <| m.succ_pos.trans_le hn), - NNReal.rpow_natCast] using hm.le.trans (pow_le_pow_right hy.le (m.le_succ.trans hn)) + NNReal.rpow_natCast] using hm.le.trans (pow_right_mono₀ hy.le (m.le_succ.trans hn)) end NNReal diff --git a/Mathlib/Analysis/SpecificLimits/Basic.lean b/Mathlib/Analysis/SpecificLimits/Basic.lean index 370cf43a6f16a..22f9e6116a899 100644 --- a/Mathlib/Analysis/SpecificLimits/Basic.lean +++ b/Mathlib/Analysis/SpecificLimits/Basic.lean @@ -119,7 +119,7 @@ theorem tendsto_natCast_div_add_atTop {𝕜 : Type*} [DivisionRing 𝕜] [Topolo theorem tendsto_add_one_pow_atTop_atTop_of_pos [LinearOrderedSemiring α] [Archimedean α] {r : α} (h : 0 < r) : Tendsto (fun n : ℕ ↦ (r + 1) ^ n) atTop atTop := - tendsto_atTop_atTop_of_monotone' (fun _ _ ↦ pow_le_pow_right <| le_add_of_nonneg_left h.le) <| + tendsto_atTop_atTop_of_monotone' (pow_right_mono₀ <| le_add_of_nonneg_left h.le) <| not_bddAbove_iff.2 fun _ ↦ Set.exists_range_iff.2 <| add_one_pow_unbounded_of_pos _ h theorem tendsto_pow_atTop_atTop_of_one_lt [LinearOrderedRing α] [Archimedean α] {r : α} @@ -513,7 +513,7 @@ theorem summable_one_div_pow_of_le {m : ℝ} {f : ℕ → ℕ} (hm : 1 < m) (fi (summable_geometric_of_lt_one (one_div_nonneg.mpr (zero_le_one.trans hm.le)) ((one_div_lt (zero_lt_one.trans hm) zero_lt_one).mpr (one_div_one.le.trans_lt hm))) rw [div_pow, one_pow] - refine (one_div_le_one_div ?_ ?_).mpr (pow_le_pow_right hm.le (fi a)) <;> + refine (one_div_le_one_div ?_ ?_).mpr (pow_right_mono₀ hm.le (fi a)) <;> exact pow_pos (zero_lt_one.trans hm) _ /-! ### Positive sequences with small sums on countable types -/ diff --git a/Mathlib/Analysis/SpecificLimits/FloorPow.lean b/Mathlib/Analysis/SpecificLimits/FloorPow.lean index 68cd1103fc6e6..e3bb53507d60d 100644 --- a/Mathlib/Analysis/SpecificLimits/FloorPow.lean +++ b/Mathlib/Analysis/SpecificLimits/FloorPow.lean @@ -226,7 +226,7 @@ theorem sum_div_pow_sq_le_div_sq (N : ℕ) {j : ℝ} (hj : 0 < j) {c : ℝ} (hc have : c ^ 3 = c ^ 2 * c := by ring simp only [mul_sub, this, mul_one, inv_pow, sub_le_sub_iff_left] rw [mul_assoc, mul_comm c, ← mul_assoc, mul_inv_cancel₀ (sq_pos_of_pos cpos).ne', one_mul] - simpa using pow_le_pow_right hc.le one_le_two + simpa using pow_right_mono₀ hc.le one_le_two have C : c⁻¹ ^ 2 < 1 := pow_lt_one₀ (inv_nonneg.2 cpos.le) (inv_lt_one hc) two_ne_zero calc (∑ i ∈ (range N).filter (j < c ^ ·), (1 : ℝ) / (c ^ i) ^ 2) ≤ @@ -270,7 +270,7 @@ theorem mul_pow_le_nat_floor_pow {c : ℝ} (hc : 1 < c) (i : ℕ) : (1 - c⁻¹) (1 - c⁻¹) * c ^ i = c ^ i - c ^ i * c⁻¹ := by ring _ ≤ c ^ i - 1 := by gcongr - simpa only [← div_eq_mul_inv, one_le_div cpos, pow_one] using le_self_pow hc.le hi + simpa only [← div_eq_mul_inv, one_le_div cpos, pow_one] using le_self_pow₀ hc.le hi _ ≤ ⌊c ^ i⌋₊ := (Nat.sub_one_lt_floor _).le /-- The sum of `1/⌊c^i⌋₊^2` above a threshold `j` is comparable to `1/j^2`, up to a multiplicative diff --git a/Mathlib/Combinatorics/SetFamily/Kleitman.lean b/Mathlib/Combinatorics/SetFamily/Kleitman.lean index 7fe6c323c2871..e534f30f51b45 100644 --- a/Mathlib/Combinatorics/SetFamily/Kleitman.lean +++ b/Mathlib/Combinatorics/SetFamily/Kleitman.lean @@ -79,5 +79,5 @@ theorem Finset.card_biUnion_le_of_intersecting (s : Finset ι) (f : ι → Finse (ih _ (fun i hi ↦ (hf₁ _ <| subset_cons _ hi).2.2) ((card_le_card <| subset_cons _).trans hs)) _).trans ?_ rw [mul_tsub, two_mul, ← pow_succ', - ← add_tsub_assoc_of_le (pow_le_pow_right' (one_le_two : (1 : ℕ) ≤ 2) tsub_le_self), + ← add_tsub_assoc_of_le (pow_right_mono₀ (one_le_two : (1 : ℕ) ≤ 2) tsub_le_self), tsub_add_eq_add_tsub hs, card_cons, add_tsub_add_eq_tsub_right] diff --git a/Mathlib/Combinatorics/SimpleGraph/Density.lean b/Mathlib/Combinatorics/SimpleGraph/Density.lean index 0c78f1e2c23b3..f71bb81713dbb 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Density.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Density.lean @@ -175,7 +175,7 @@ theorem edgeDensity_sub_edgeDensity_le_one_sub_mul (hs : s₂ ⊆ s₁) (ht : t refine (sub_le_sub_left (mul_edgeDensity_le_edgeDensity r hs ht hs₂ ht₂) _).trans ?_ refine le_trans ?_ (mul_le_of_le_one_right ?_ (edgeDensity_le_one r s₂ t₂)) · rw [sub_mul, one_mul] - refine sub_nonneg_of_le (mul_le_one ?_ ?_ ?_) + refine sub_nonneg_of_le (mul_le_one₀ ?_ ?_ ?_) · exact div_le_one_of_le₀ ((@Nat.cast_le ℚ).2 (card_le_card hs)) (Nat.cast_nonneg _) · apply div_nonneg <;> exact mod_cast Nat.zero_le _ · exact div_le_one_of_le₀ ((@Nat.cast_le ℚ).2 (card_le_card ht)) (Nat.cast_nonneg _) diff --git a/Mathlib/Combinatorics/SimpleGraph/Regularity/Chunk.lean b/Mathlib/Combinatorics/SimpleGraph/Regularity/Chunk.lean index d903c4c522373..e9ac95f5e3f00 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Regularity/Chunk.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Regularity/Chunk.lean @@ -114,7 +114,7 @@ private theorem card_nonuniformWitness_sdiff_biUnion_star (hV : V ∈ P.parts) ( rw [sum_const] refine mul_le_mul_right' ?_ _ have t := card_filter_atomise_le_two_pow (s := U) hX - refine t.trans (pow_le_pow_right (by norm_num) <| tsub_le_tsub_right ?_ _) + refine t.trans (pow_right_mono₀ (by norm_num) <| tsub_le_tsub_right ?_ _) exact card_image_le.trans (card_le_card <| filter_subset _ _) private theorem one_sub_eps_mul_card_nonuniformWitness_le_card_star (hV : V ∈ P.parts) diff --git a/Mathlib/Combinatorics/SimpleGraph/Regularity/Lemma.lean b/Mathlib/Combinatorics/SimpleGraph/Regularity/Lemma.lean index 5e120205aeeb8..bee7ea6b54e48 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Regularity/Lemma.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Regularity/Lemma.lean @@ -128,7 +128,7 @@ theorem szemeredi_regularity (hε : 0 < ε) (hl : l ≤ card α) : -- We gather a few numerical facts. have hεl' : 100 ≤ 4 ^ P.parts.card * ε ^ 5 := (hundred_lt_pow_initialBound_mul hε l).le.trans - (mul_le_mul_of_nonneg_right (pow_le_pow_right (by norm_num) hP₂) <| by positivity) + (mul_le_mul_of_nonneg_right (pow_right_mono₀ (by norm_num) hP₂) <| by positivity) have hi : (i : ℝ) ≤ 4 / ε ^ 5 := by have hi : ε ^ 5 / 4 * ↑i ≤ 1 := hP₄.trans (mod_cast P.energy_le_one G) rw [div_mul_eq_mul_div, div_le_iff₀ (show (0 : ℝ) < 4 by norm_num)] at hi diff --git a/Mathlib/Computability/Ackermann.lean b/Mathlib/Computability/Ackermann.lean index 7017d99c1eb68..211e594b14670 100644 --- a/Mathlib/Computability/Ackermann.lean +++ b/Mathlib/Computability/Ackermann.lean @@ -3,7 +3,6 @@ Copyright (c) 2022 Violeta Hernández Palacios. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Violeta Hernández Palacios -/ -import Mathlib.Algebra.Order.Ring.Basic import Mathlib.Computability.Primrec import Mathlib.Tactic.Ring import Mathlib.Tactic.Linarith @@ -94,7 +93,7 @@ theorem ack_three (n : ℕ) : ack 3 n = 2 ^ (n + 3) - 3 := by have H : 2 * 3 ≤ 2 * 2 ^ 3 := by norm_num apply H.trans rw [_root_.mul_le_mul_left two_pos] - exact pow_le_pow_right one_le_two (Nat.le_add_left 3 n) + exact pow_right_mono₀ one_le_two (Nat.le_add_left 3 n) theorem ack_pos : ∀ m n, 0 < ack m n | 0, n => by simp diff --git a/Mathlib/Data/Complex/Exponential.lean b/Mathlib/Data/Complex/Exponential.lean index 9c2622348ebe2..d6a37e2bade95 100644 --- a/Mathlib/Data/Complex/Exponential.lean +++ b/Mathlib/Data/Complex/Exponential.lean @@ -1336,7 +1336,7 @@ theorem cos_pos_of_le_one {x : ℝ} (hx : |x| ≤ 1) : 0 < cos x := gcongr · exact pow_le_one₀ (abs_nonneg _) hx · rw [sq, ← abs_mul_self, abs_mul] - exact mul_le_one hx (abs_nonneg _) hx + exact mul_le_one₀ hx (abs_nonneg _) hx _ < 1 := by norm_num) _ ≤ cos x := sub_le_comm.1 (abs_sub_le_iff.1 (cos_bound hx)).2 diff --git a/Mathlib/Data/ENNReal/Inv.lean b/Mathlib/Data/ENNReal/Inv.lean index 5847cba382776..1fc1f2eb0733b 100644 --- a/Mathlib/Data/ENNReal/Inv.lean +++ b/Mathlib/Data/ENNReal/Inv.lean @@ -574,13 +574,13 @@ theorem Ioo_zero_top_eq_iUnion_Ico_zpow {y : ℝ≥0∞} (hy : 1 < y) (h'y : y theorem zpow_le_of_le {x : ℝ≥0∞} (hx : 1 ≤ x) {a b : ℤ} (h : a ≤ b) : x ^ a ≤ x ^ b := by induction' a with a a <;> induction' b with b b · simp only [Int.ofNat_eq_coe, zpow_natCast] - exact pow_le_pow_right hx (Int.le_of_ofNat_le_ofNat h) + exact pow_right_mono₀ hx (Int.le_of_ofNat_le_ofNat h) · apply absurd h (not_le_of_gt _) exact lt_of_lt_of_le (Int.negSucc_lt_zero _) (Int.ofNat_nonneg _) · simp only [zpow_negSucc, Int.ofNat_eq_coe, zpow_natCast] refine (ENNReal.inv_le_one.2 ?_).trans ?_ <;> exact one_le_pow_of_one_le' hx _ · simp only [zpow_negSucc, ENNReal.inv_le_inv] - apply pow_le_pow_right hx + apply pow_right_mono₀ hx simpa only [← Int.ofNat_le, neg_le_neg_iff, Int.ofNat_add, Int.ofNat_one, Int.negSucc_eq] using h diff --git a/Mathlib/Data/Nat/Choose/Factorization.lean b/Mathlib/Data/Nat/Choose/Factorization.lean index 71966a049eb3a..258c4eb897817 100644 --- a/Mathlib/Data/Nat/Choose/Factorization.lean +++ b/Mathlib/Data/Nat/Choose/Factorization.lean @@ -78,7 +78,7 @@ theorem factorization_choose_of_lt_three_mul (hp' : p ≠ 2) (hk : p ≤ k) (hk' n < 3 * p := hn _ ≤ p * p := mul_le_mul_right' this p _ = p ^ 2 := (sq p).symm - _ ≤ p ^ i := pow_le_pow_right hp.one_lt.le hi + _ ≤ p ^ i := pow_right_mono₀ hp.one_lt.le hi rwa [mod_eq_of_lt (lt_of_le_of_lt hkn hn), mod_eq_of_lt (lt_of_le_of_lt tsub_le_self hn), add_tsub_cancel_of_le hkn] diff --git a/Mathlib/Data/Real/Pi/Bounds.lean b/Mathlib/Data/Real/Pi/Bounds.lean index 9a77915281ee6..131af1b4136f2 100644 --- a/Mathlib/Data/Real/Pi/Bounds.lean +++ b/Mathlib/Data/Real/Pi/Bounds.lean @@ -40,7 +40,7 @@ theorem pi_lt_sqrtTwoAddSeries (n : ℕ) : · rw [div_le_iff₀'] · refine le_trans pi_le_four ?_ simp only [show (4 : ℝ) = (2 : ℝ) ^ 2 by norm_num, mul_one] - apply pow_le_pow_right (by norm_num) + apply pow_right_mono₀ (by norm_num) apply le_add_of_nonneg_left; apply Nat.zero_le · apply pow_pos; norm_num apply add_le_add_left; rw [div_le_div_right (by norm_num)] diff --git a/Mathlib/Data/Real/Pi/Irrational.lean b/Mathlib/Data/Real/Pi/Irrational.lean index 7301ad0ffca5c..e3884b774698d 100644 --- a/Mathlib/Data/Real/Pi/Irrational.lean +++ b/Mathlib/Data/Real/Pi/Irrational.lean @@ -251,7 +251,7 @@ private lemma I_le (n : ℕ) : I n (π / 2) ≤ 2 := by intros x hx simp only [uIoc_of_le, neg_le_self_iff, zero_le_one, mem_Ioc] at hx rw [norm_eq_abs, abs_mul, abs_pow] - refine mul_le_one (pow_le_one₀ (abs_nonneg _) ?_) (abs_nonneg _) (abs_cos_le_one _) + refine mul_le_one₀ (pow_le_one₀ (abs_nonneg _) ?_) (abs_nonneg _) (abs_cos_le_one _) rw [abs_le] constructor <;> nlinarith diff --git a/Mathlib/NumberTheory/Bertrand.lean b/Mathlib/NumberTheory/Bertrand.lean index 772c5d12f2dff..f35408c70ebe0 100644 --- a/Mathlib/NumberTheory/Bertrand.lean +++ b/Mathlib/NumberTheory/Bertrand.lean @@ -165,13 +165,13 @@ theorem centralBinom_le_of_no_bertrand_prime (n : ℕ) (n_large : 2 < n) · exact pow_factorization_choose_le (mul_pos two_pos n_pos) have : (Finset.Icc 1 (sqrt (2 * n))).card = sqrt (2 * n) := by rw [card_Icc, Nat.add_sub_cancel] rw [Finset.prod_const] - refine pow_le_pow_right n2_pos ((Finset.card_le_card fun x hx => ?_).trans this.le) + refine pow_right_mono₀ n2_pos ((Finset.card_le_card fun x hx => ?_).trans this.le) obtain ⟨h1, h2⟩ := Finset.mem_filter.1 hx exact Finset.mem_Icc.mpr ⟨(Finset.mem_filter.1 h1).2.one_lt.le, h2⟩ · refine le_trans ?_ (primorial_le_4_pow (2 * n / 3)) refine (Finset.prod_le_prod' fun p hp => (?_ : f p ≤ p)).trans ?_ · obtain ⟨h1, h2⟩ := Finset.mem_filter.1 hp - refine (pow_le_pow_right (Finset.mem_filter.1 h1).2.one_lt.le ?_).trans (pow_one p).le + refine (pow_right_mono₀ (Finset.mem_filter.1 h1).2.one_lt.le ?_).trans (pow_one p).le exact Nat.factorization_choose_le_one (sqrt_lt'.mp <| not_le.1 h2) refine Finset.prod_le_prod_of_subset_of_one_le' (Finset.filter_subset _ _) ?_ exact fun p hp _ => (Finset.mem_filter.1 hp).2.one_lt.le diff --git a/Mathlib/NumberTheory/ClassNumber/AdmissibleCardPowDegree.lean b/Mathlib/NumberTheory/ClassNumber/AdmissibleCardPowDegree.lean index c1fd34862bfb8..fabb26655f44b 100644 --- a/Mathlib/NumberTheory/ClassNumber/AdmissibleCardPowDegree.lean +++ b/Mathlib/NumberTheory/ClassNumber/AdmissibleCardPowDegree.lean @@ -160,7 +160,7 @@ theorem cardPowDegree_anti_archimedean {x y z : Fq[X]} {a : ℤ} (hxy : cardPowD cardPowDegree_nonzero _ hyz'] have : (1 : ℤ) ≤ Fintype.card Fq := mod_cast (@Fintype.one_lt_card Fq _ _).le simp only [Int.cast_pow, Int.cast_natCast, le_max_iff] - refine Or.imp (pow_le_pow_right this) (pow_le_pow_right this) ?_ + refine Or.imp (pow_le_pow_right₀ this) (pow_le_pow_right₀ this) ?_ rw [natDegree_le_iff_degree_le, natDegree_le_iff_degree_le, ← le_max_iff, ← degree_eq_natDegree hxy', ← degree_eq_natDegree hyz'] convert degree_add_le (x - y) (y - z) using 2 diff --git a/Mathlib/NumberTheory/FermatPsp.lean b/Mathlib/NumberTheory/FermatPsp.lean index f077508bbee07..113b74cc7f237 100644 --- a/Mathlib/NumberTheory/FermatPsp.lean +++ b/Mathlib/NumberTheory/FermatPsp.lean @@ -135,7 +135,7 @@ private theorem b_id_helper {a b : ℕ} (ha : 2 ≤ a) (hb : 2 < b) : 2 ≤ (a ^ calc 2 * a + 1 ≤ a ^ 2 * a := by nlinarith _ = a ^ 3 := by rw [Nat.pow_succ a 2] - _ ≤ a ^ b := pow_le_pow_right (Nat.le_of_succ_le ha) hb + _ ≤ a ^ b := pow_right_mono₀ (Nat.le_of_succ_le ha) hb private theorem AB_id_helper (b p : ℕ) (_ : 2 ≤ b) (hp : Odd p) : (b ^ p - 1) / (b - 1) * ((b ^ p + 1) / (b + 1)) = (b ^ (2 * p) - 1) / (b ^ 2 - 1) := by diff --git a/Mathlib/NumberTheory/Liouville/LiouvilleNumber.lean b/Mathlib/NumberTheory/Liouville/LiouvilleNumber.lean index 3fb42ea6a0f27..49a8325713112 100644 --- a/Mathlib/NumberTheory/Liouville/LiouvilleNumber.lean +++ b/Mathlib/NumberTheory/Liouville/LiouvilleNumber.lean @@ -145,7 +145,7 @@ theorem aux_calc (n : ℕ) {m : ℝ} (hm : 2 ≤ m) : any_goals exact pow_pos (zero_lt_two.trans_le hm) _ -- `2 ≤ m ^ n!` is a consequence of monotonicity of exponentiation at `2 ≤ m`. exact _root_.trans (_root_.trans hm (pow_one _).symm.le) - (pow_right_mono (one_le_two.trans hm) n.factorial_pos) + (pow_right_mono₀ (one_le_two.trans hm) n.factorial_pos) _ = 1 / (m ^ n !) ^ n := congr_arg (1 / ·) (pow_mul m n ! n) /-- An upper estimate on the remainder. This estimate works with `m ∈ ℝ` satisfying `2 ≤ m` and is diff --git a/Mathlib/NumberTheory/Padics/Hensel.lean b/Mathlib/NumberTheory/Padics/Hensel.lean index 0de565647f9ee..05f97027ce268 100644 --- a/Mathlib/NumberTheory/Padics/Hensel.lean +++ b/Mathlib/NumberTheory/Padics/Hensel.lean @@ -156,7 +156,7 @@ private theorem calc_norm_le_one {n : ℕ} {z : ℤ_[p]} (hz : ih n z) : gcongr apply hz.2 _ = ‖F.derivative.eval a‖ * T ^ 2 ^ n := div_sq_cancel _ _ - _ ≤ 1 := mul_le_one (PadicInt.norm_le_one _) (T_pow_nonneg _) (le_of_lt (T_pow' hnorm _)) + _ ≤ 1 := mul_le_one₀ (PadicInt.norm_le_one _) (T_pow_nonneg _) (le_of_lt (T_pow' hnorm _)) private theorem calc_deriv_dist {z z' z1 : ℤ_[p]} (hz' : z' = z - z1) @@ -183,7 +183,7 @@ private def calc_eval_z' {z z' z1 : ℤ_[p]} (hz' : z' = z - z1) {n} (hz : ih n obtain ⟨q, hq⟩ := F.binomExpansion z (-z1) have : ‖(↑(F.derivative.eval z) * (↑(F.eval z) / ↑(F.derivative.eval z)) : ℚ_[p])‖ ≤ 1 := by rw [padicNormE.mul] - exact mul_le_one (PadicInt.norm_le_one _) (norm_nonneg _) h1 + exact mul_le_one₀ (PadicInt.norm_le_one _) (norm_nonneg _) h1 have : F.derivative.eval z * -z1 = -F.eval z := by calc F.derivative.eval z * -z1 = @@ -275,7 +275,7 @@ private theorem newton_seq_dist_aux (n : ℕ) : | 0 => by simp [T_pow_nonneg, mul_nonneg] | k + 1 => have : 2 ^ n ≤ 2 ^ (n + k) := by - apply pow_le_pow_right + apply pow_right_mono₀ · norm_num · apply Nat.le_add_right calc @@ -356,7 +356,7 @@ private theorem T_pos : T > 0 := by private theorem newton_seq_succ_dist_weak (n : ℕ) : ‖newton_seq (n + 2) - newton_seq (n + 1)‖ < ‖F.eval a‖ / ‖F.derivative.eval a‖ := have : 2 ≤ 2 ^ (n + 1) := by - have := pow_le_pow_right (by norm_num : 1 ≤ 2) (Nat.le_add_left _ _ : 1 ≤ n + 1) + have := pow_right_mono₀ (by norm_num : 1 ≤ 2) (Nat.le_add_left _ _ : 1 ≤ n + 1) simpa using this calc ‖newton_seq (n + 2) - newton_seq (n + 1)‖ ≤ ‖F.derivative.eval a‖ * T ^ 2 ^ (n + 1) := diff --git a/Mathlib/NumberTheory/Padics/PadicIntegers.lean b/Mathlib/NumberTheory/Padics/PadicIntegers.lean index 02820e33e2483..06fae1357239d 100644 --- a/Mathlib/NumberTheory/Padics/PadicIntegers.lean +++ b/Mathlib/NumberTheory/Padics/PadicIntegers.lean @@ -78,7 +78,7 @@ def subring : Subring ℚ_[p] where zero_mem' := by norm_num one_mem' := by norm_num add_mem' hx hy := (padicNormE.nonarchimedean _ _).trans <| max_le_iff.2 ⟨hx, hy⟩ - mul_mem' hx hy := (padicNormE.mul _ _).trans_le <| mul_le_one hx (norm_nonneg _) hy + mul_mem' hx hy := (padicNormE.mul _ _).trans_le <| mul_le_one₀ hx (norm_nonneg _) hy neg_mem' hx := (norm_neg _).trans_le hx @[simp] diff --git a/Mathlib/Order/Filter/AtTopBot/Ring.lean b/Mathlib/Order/Filter/AtTopBot/Ring.lean index 874630711de2c..74207b0e55f14 100644 --- a/Mathlib/Order/Filter/AtTopBot/Ring.lean +++ b/Mathlib/Order/Filter/AtTopBot/Ring.lean @@ -3,8 +3,8 @@ Copyright (c) 2019 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ +import Mathlib.Algebra.Order.Ring.Defs import Mathlib.Order.Filter.AtTopBot.Group -import Mathlib.Algebra.Order.Ring.Basic /-! # Convergence to ±infinity in ordered rings @@ -30,7 +30,7 @@ theorem tendsto_mul_self_atTop : Tendsto (fun x : α => x * x) atTop atTop := /-- The monomial function `x^n` tends to `+∞` at `+∞` for any positive natural `n`. A version for positive real powers exists as `tendsto_rpow_atTop`. -/ theorem tendsto_pow_atTop {n : ℕ} (hn : n ≠ 0) : Tendsto (fun x : α => x ^ n) atTop atTop := - tendsto_atTop_mono' _ ((eventually_ge_atTop 1).mono fun _x hx => le_self_pow hx hn) tendsto_id + tendsto_atTop_mono' _ ((eventually_ge_atTop 1).mono fun _x hx => le_self_pow₀ hx hn) tendsto_id end OrderedSemiring diff --git a/Mathlib/Probability/StrongLaw.lean b/Mathlib/Probability/StrongLaw.lean index 94bcc41f1f6db..34d9233a9272f 100644 --- a/Mathlib/Probability/StrongLaw.lean +++ b/Mathlib/Probability/StrongLaw.lean @@ -396,7 +396,7 @@ theorem strong_law_aux1 {c : ℝ} (c_one : 1 < c) {ε : ℝ} (εpos : 0 < ε) : set Y := fun n : ℕ => truncation (X n) n set S := fun n => ∑ i ∈ range n, Y i with hS let u : ℕ → ℕ := fun n => ⌊c ^ n⌋₊ - have u_mono : Monotone u := fun i j hij => Nat.floor_mono (pow_le_pow_right c_one.le hij) + have u_mono : Monotone u := fun i j hij => Nat.floor_mono (pow_right_mono₀ c_one.le hij) have I1 : ∀ K, ∑ j ∈ range K, ((j : ℝ) ^ 2)⁻¹ * Var[Y j] ≤ 2 * 𝔼[X 0] := by intro K calc diff --git a/Mathlib/Topology/Algebra/PontryaginDual.lean b/Mathlib/Topology/Algebra/PontryaginDual.lean index fe184b35a5fa5..262a7e2fde965 100644 --- a/Mathlib/Topology/Algebra/PontryaginDual.lean +++ b/Mathlib/Topology/Algebra/PontryaginDual.lean @@ -64,7 +64,7 @@ instance [LocallyCompactSpace G] : LocallyCompactSpace (PontryaginDual G) := by apply Set.Ioo_subset_Ioc_self rw [← two_mul, Set.mem_Ioo, ← abs_lt, abs_mul, abs_two, ← lt_div_iff₀' two_pos] exact h1.trans_le - (div_le_div_of_nonneg_left Real.pi_nonneg two_pos (le_self_pow one_le_two n.succ_ne_zero)) + (div_le_div_of_nonneg_left Real.pi_nonneg two_pos (le_self_pow₀ one_le_two n.succ_ne_zero)) · rw [← Circle.exp_zero, ← isLocalHomeomorph_circleExp.map_nhds_eq 0] refine ((nhds_basis_zero_abs_sub_lt ℝ).to_hasBasis (fun x hx ↦ ⟨Nat.ceil (Real.pi / x), trivial, fun t ht ↦ ?_⟩) diff --git a/Mathlib/Topology/UnitInterval.lean b/Mathlib/Topology/UnitInterval.lean index c1aff48fbc13d..55c6405dda1b6 100644 --- a/Mathlib/Topology/UnitInterval.lean +++ b/Mathlib/Topology/UnitInterval.lean @@ -40,7 +40,7 @@ theorem one_mem : (1 : ℝ) ∈ I := ⟨zero_le_one, le_rfl⟩ theorem mul_mem {x y : ℝ} (hx : x ∈ I) (hy : y ∈ I) : x * y ∈ I := - ⟨mul_nonneg hx.1 hy.1, mul_le_one hx.2 hy.1 hy.2⟩ + ⟨mul_nonneg hx.1 hy.1, mul_le_one₀ hx.2 hy.1 hy.2⟩ theorem div_mem {x y : ℝ} (hx : 0 ≤ x) (hy : 0 ≤ y) (hxy : x ≤ y) : x / y ∈ I := ⟨div_nonneg hx hy, div_le_one_of_le₀ hxy hy⟩ From 4d57c2c097b1de2f0a00fa645c49c848e2e19f23 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Mon, 7 Oct 2024 17:59:18 +0000 Subject: [PATCH 32/43] feat: Noetherian/Artinian is closed under extensions (submodule form) (#17425) + Add `isNoetherian/Artinian_iff_submodule_quotient`. + Remove unnecessary assumptions from `isArtinian_of_range_eq_ker`, mimicking #12453. + Move up two instances `isArtinian_of_quotient_of_artinian` and `isNoetherian_of_finite`. + Golf `isNoetherian/Artinian_pi`. --- Mathlib/RingTheory/Artinian.lean | 44 +++++++------ Mathlib/RingTheory/Noetherian.lean | 100 ++++++----------------------- 2 files changed, 43 insertions(+), 101 deletions(-) diff --git a/Mathlib/RingTheory/Artinian.lean b/Mathlib/RingTheory/Artinian.lean index 60965f6c689e4..e3f2bcd132e70 100644 --- a/Mathlib/RingTheory/Artinian.lean +++ b/Mathlib/RingTheory/Artinian.lean @@ -91,22 +91,33 @@ 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} theorem isArtinian_of_linearEquiv (f : M ≃ₗ[R] P) [IsArtinian R M] : IsArtinian R P := isArtinian_of_surjective _ f.toLinearMap f.toEquiv.surjective theorem isArtinian_of_range_eq_ker [IsArtinian R M] [IsArtinian R P] (f : M →ₗ[R] N) (g : N →ₗ[R] P) - (hf : Function.Injective f) (hg : Function.Surjective g) (h : LinearMap.range f = LinearMap.ker g) : IsArtinian R N := - wellFounded_lt_exact_sequence (LinearMap.range f) (Submodule.map f) - (Submodule.comap f) (Submodule.comap g) (Submodule.map g) (Submodule.gciMapComap hf) - (Submodule.giMapComap hg) - (by simp [Submodule.map_comap_eq, inf_comm]) (by simp [Submodule.comap_map_eq, h]) + 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.inl_injective - LinearMap.snd_surjective (LinearMap.range_inl 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 _⟩ @@ -115,17 +126,12 @@ instance (priority := 100) isArtinian_of_finite [Finite M] : IsArtinian R M := -- attribute [local elab_as_elim] Finite.induction_empty_option instance isArtinian_pi {R ι : Type*} [Finite ι] : - ∀ {M : ι → Type*} [Ring R] [∀ i, AddCommGroup (M i)], - ∀ [∀ i, Module R (M i)], ∀ [∀ i, IsArtinian R (M i)], IsArtinian R (∀ i, M i) := by + ∀ {M : ι → Type*} [Ring R] [∀ 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 _ _ _ ι - · intro α β e hα M _ _ _ _ - have := @hα - exact isArtinian_of_linearEquiv (LinearEquiv.piCongrLeft R M e) - · intro M _ _ _ _ - infer_instance - · intro α _ ih M _ _ _ _ - have := @ih - exact isArtinian_of_linearEquiv (LinearEquiv.piOptionEquivProd R).symm + · 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 @@ -336,10 +342,6 @@ theorem Ring.isArtinian_of_zero_eq_one {R} [Ring R] (h01 : (0 : R) = 1) : IsArti theorem isArtinian_of_submodule_of_artinian (R M) [Ring R] [AddCommGroup M] [Module R M] (N : Submodule R M) (_ : IsArtinian R M) : IsArtinian R N := inferInstance -instance isArtinian_of_quotient_of_artinian (R) [Ring R] (M) [AddCommGroup M] [Module R M] - (N : Submodule R M) [IsArtinian R M] : IsArtinian R (M ⧸ N) := - isArtinian_of_surjective M (Submodule.mkQ N) (Submodule.Quotient.mk_surjective N) - /-- 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] [Module S M] [Module R M] [IsScalarTower R S M] (h : IsArtinian R M) : IsArtinian S M := diff --git a/Mathlib/RingTheory/Noetherian.lean b/Mathlib/RingTheory/Noetherian.lean index d89b6584db58f..62188f744de54 100644 --- a/Mathlib/RingTheory/Noetherian.lean +++ b/Mathlib/RingTheory/Noetherian.lean @@ -144,6 +144,10 @@ variable {R M N : Type*} variable [Semiring R] [AddCommMonoid M] [AddCommMonoid N] [Module R M] [Module R N] variable (R M) +-- see Note [lower instance priority] +instance (priority := 80) _root_.isNoetherian_of_finite [Finite M] : IsNoetherian R M := + ⟨fun s => ⟨(s : Set M).toFinite.toFinset, by rw [Set.Finite.coe_toFinset, Submodule.span_eq]⟩⟩ + -- see Note [lower instance priority] instance (priority := 100) IsNoetherian.finite [IsNoetherian R M] : Module.Finite R M := ⟨IsNoetherian.noetherian ⊤⟩ @@ -184,77 +188,13 @@ instance isNoetherian_prod [IsNoetherian R M] [IsNoetherian R P] : IsNoetherian fun x ⟨_, hx2⟩ => ⟨x.1, Prod.ext rfl <| Eq.symm <| LinearMap.mem_ker.1 hx2⟩ Submodule.map_comap_eq_self this ▸ (noetherian _).map _⟩ -instance isNoetherian_pi {R ι : Type*} {M : ι → Type*} - [Ring R] [∀ i, AddCommGroup (M i)] [∀ i, Module R (M i)] [Finite ι] - [∀ i, IsNoetherian R (M i)] : IsNoetherian R (∀ i, M i) := by - cases nonempty_fintype ι - haveI := Classical.decEq ι - suffices on_finset : ∀ s : Finset ι, IsNoetherian R (∀ i : s, M i) by - let coe_e := Equiv.subtypeUnivEquiv <| @Finset.mem_univ ι _ - letI : IsNoetherian R (∀ i : Finset.univ, M (coe_e i)) := on_finset Finset.univ - exact isNoetherian_of_linearEquiv (LinearEquiv.piCongrLeft R M coe_e) - intro s - induction' s using Finset.induction with a s has ih - · exact ⟨fun s => by - have : s = ⊥ := by simp only [eq_iff_true_of_subsingleton] - rw [this] - apply Submodule.fg_bot⟩ - refine - @isNoetherian_of_linearEquiv R (M a × ((i : s) → M i)) _ _ _ _ _ _ ?_ <| - @isNoetherian_prod R (M a) _ _ _ _ _ _ _ ih - refine - { toFun := fun f i => - (Finset.mem_insert.1 i.2).by_cases - (fun h : i.1 = a => show M i.1 from Eq.recOn h.symm f.1) - (fun h : i.1 ∈ s => show M i.1 from f.2 ⟨i.1, h⟩), - invFun := fun f => - (f ⟨a, Finset.mem_insert_self _ _⟩, fun i => f ⟨i.1, Finset.mem_insert_of_mem i.2⟩), - map_add' := ?_, - map_smul' := ?_ - left_inv := ?_, - right_inv := ?_ } - · intro f g - ext i - unfold Or.by_cases - cases' i with i hi - rcases Finset.mem_insert.1 hi with (rfl | h) - · change _ = _ + _ - simp only [dif_pos] - rfl - · change _ = _ + _ - have : ¬i = a := by - rintro rfl - exact has h - simp only [dif_neg this, dif_pos h] - rfl - · intro c f - ext i - unfold Or.by_cases - cases' i with i hi - rcases Finset.mem_insert.1 hi with (rfl | h) - · dsimp - simp only [dif_pos] - · dsimp - have : ¬i = a := by - rintro rfl - exact has h - simp only [dif_neg this, dif_pos h] - · intro f - apply Prod.ext - · simp only [Or.by_cases, dif_pos] - · ext ⟨i, his⟩ - have : ¬i = a := by - rintro rfl - exact has his - simp only [Or.by_cases, this, not_false_iff, dif_neg] - · intro f - ext ⟨i, hi⟩ - rcases Finset.mem_insert.1 hi with (rfl | h) - · simp only [Or.by_cases, dif_pos] - · have : ¬i = a := by - rintro rfl - exact has h - simp only [Or.by_cases, dif_neg this, dif_pos h] +instance isNoetherian_pi {R ι : Type*} [Finite ι] : + ∀ {M : ι → Type*} [Ring R] [∀ i, AddCommGroup (M i)] + [∀ 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 + · exact fun ih ↦ isNoetherian_of_linearEquiv (LinearEquiv.piOptionEquivProd R).symm /-- A version of `isNoetherian_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 @@ -389,15 +329,20 @@ theorem isNoetherian_of_range_eq_ker [IsNoetherian R P] 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.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_refl _)) + (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] + /-- For an endomorphism of a Noetherian module, any sufficiently large iterate has disjoint kernel and range. -/ theorem LinearMap.eventually_disjoint_ker_pow_range_pow (f : M →ₗ[R] M) : @@ -500,11 +445,6 @@ theorem isNoetherianRing_iff_ideal_fg (R : Type*) [Semiring R] : IsNoetherianRing R ↔ ∀ I : Ideal R, I.FG := isNoetherianRing_iff.trans isNoetherian_def --- see Note [lower instance priority] -instance (priority := 80) isNoetherian_of_finite (R M) [Finite M] [Semiring R] [AddCommMonoid M] - [Module R M] : IsNoetherian R M := - ⟨fun s => ⟨(s : Set M).toFinite.toFinset, by rw [Set.Finite.coe_toFinset, Submodule.span_eq]⟩⟩ - -- see Note [lower instance priority] /-- Modules over the trivial ring are Noetherian. -/ instance (priority := 100) isNoetherian_of_subsingleton (R M) [Subsingleton R] [Semiring R] From 5151ec1ccf5c2abd8165b9553c0bc6cb58fe07d9 Mon Sep 17 00:00:00 2001 From: Yongle Hu Date: Mon, 7 Oct 2024 17:59:19 +0000 Subject: [PATCH 33/43] feat(RingTheory/Ideal/Operations): drop `IsDedekindDomain` assumptions of some lemmas and move them to `Ideal/Operations.lean` (#17460) Drop `[IsDedekindDomain R]` assumptions of some lemmas about ideal pow le a prime ideal and move them to `RingTheory/Ideal/Operations.lean`. Also add a lemma about product of elements in a prime ideal. Co-authored-by: Hu Yongle <2065545849@qq.com> Co-authored-by: Yongle Hu <2065545849@qq.com> --- Mathlib/RingTheory/DedekindDomain/Ideal.lean | 40 ++++---------------- Mathlib/RingTheory/Ideal/Operations.lean | 27 +++++++++++++ 2 files changed, 34 insertions(+), 33 deletions(-) diff --git a/Mathlib/RingTheory/DedekindDomain/Ideal.lean b/Mathlib/RingTheory/DedekindDomain/Ideal.lean index 0c6332b7d5a5b..426d939612987 100644 --- a/Mathlib/RingTheory/DedekindDomain/Ideal.lean +++ b/Mathlib/RingTheory/DedekindDomain/Ideal.lean @@ -1191,26 +1191,6 @@ theorem Ideal.le_mul_of_no_prime_factors {I J K : Ideal R} (UniqueFactorizationMonoid.dvd_of_dvd_mul_right_of_no_prime_factors (b := K) hJ0 ?_ hJ) exact fun hPJ hPK => mt Ideal.isPrime_of_prime (coprime _ hPJ hPK) -theorem Ideal.le_of_pow_le_prime {I P : Ideal R} [hP : P.IsPrime] {n : ℕ} (h : I ^ n ≤ P) : - I ≤ P := by - by_cases hP0 : P = ⊥ - · simp only [hP0, le_bot_iff] at h ⊢ - exact pow_eq_zero h - rw [← Ideal.dvd_iff_le] at h ⊢ - exact ((Ideal.prime_iff_isPrime hP0).mpr hP).dvd_of_dvd_pow h - -theorem Ideal.pow_le_prime_iff {I P : Ideal R} [_hP : P.IsPrime] {n : ℕ} (hn : n ≠ 0) : - I ^ n ≤ P ↔ I ≤ P := - ⟨Ideal.le_of_pow_le_prime, fun h => _root_.trans (Ideal.pow_le_self hn) h⟩ - -theorem Ideal.prod_le_prime {ι : Type*} {s : Finset ι} {f : ι → Ideal R} {P : Ideal R} - [hP : P.IsPrime] : ∏ i ∈ s, f i ≤ P ↔ ∃ i ∈ s, f i ≤ P := by - by_cases hP0 : P = ⊥ - · simp only [hP0, le_bot_iff] - rw [← Ideal.zero_eq_bot, Finset.prod_eq_zero_iff] - simp only [← Ideal.dvd_iff_le] - exact ((Ideal.prime_iff_isPrime hP0).mpr hP).dvd_finset_prod_iff _ - /-- The intersection of distinct prime powers in a Dedekind domain is the product of these prime powers. -/ theorem IsDedekindDomain.inf_prime_pow_eq_prod {ι : Type*} (s : Finset ι) (f : ι → Ideal R) @@ -1228,15 +1208,13 @@ theorem IsDedekindDomain.inf_prime_pow_eq_prod {ι : Type*} (s : Finset ι) (f : rw [Finset.inf_insert, Finset.prod_insert ha, ih] refine le_antisymm (Ideal.le_mul_of_no_prime_factors ?_ inf_le_left inf_le_right) Ideal.mul_le_inf intro P hPa hPs hPp - obtain ⟨b, hb, hPb⟩ := Ideal.prod_le_prime.mp hPs + obtain ⟨b, hb, hPb⟩ := hPp.prod_le.mp hPs haveI := Ideal.isPrime_of_prime (prime a (Finset.mem_insert_self a s)) haveI := Ideal.isPrime_of_prime (prime b (Finset.mem_insert_of_mem hb)) refine coprime a (Finset.mem_insert_self a s) b (Finset.mem_insert_of_mem hb) ?_ ?_ · exact (ne_of_mem_of_not_mem hb ha).symm - · refine ((Ring.DimensionLeOne.prime_le_prime_iff_eq ?_).mp - (Ideal.le_of_pow_le_prime hPa)).trans - ((Ring.DimensionLeOne.prime_le_prime_iff_eq ?_).mp - (Ideal.le_of_pow_le_prime hPb)).symm + · refine ((Ring.DimensionLeOne.prime_le_prime_iff_eq ?_).mp (hPp.le_of_pow_le hPa)).trans + ((Ring.DimensionLeOne.prime_le_prime_iff_eq ?_).mp (hPp.le_of_pow_le hPb)).symm · exact (prime a (Finset.mem_insert_self a s)).ne_zero · exact (prime b (Finset.mem_insert_of_mem hb)).ne_zero @@ -1251,17 +1229,13 @@ noncomputable def IsDedekindDomain.quotientEquivPiOfProdEq {ι : Type*} [Fintype simp only [← prod_eq, Finset.inf_eq_iInf, Finset.mem_univ, ciInf_pos, ← IsDedekindDomain.inf_prime_pow_eq_prod _ _ _ (fun i _ => prime i) (coprime.set_pairwise _)])).trans <| - Ideal.quotientInfRingEquivPiQuotient _ fun i j hij => Ideal.coprime_of_no_prime_ge (by + Ideal.quotientInfRingEquivPiQuotient _ fun i j hij => Ideal.coprime_of_no_prime_ge <| by intro P hPi hPj hPp haveI := Ideal.isPrime_of_prime (prime i) haveI := Ideal.isPrime_of_prime (prime j) - refine coprime hij ?_ - refine ((Ring.DimensionLeOne.prime_le_prime_iff_eq ?_).mp - (Ideal.le_of_pow_le_prime hPi)).trans - ((Ring.DimensionLeOne.prime_le_prime_iff_eq ?_).mp - (Ideal.le_of_pow_le_prime hPj)).symm - · exact (prime i).ne_zero - · exact (prime j).ne_zero) + exact coprime hij <| ((Ring.DimensionLeOne.prime_le_prime_iff_eq (prime i).ne_zero).mp + (hPp.le_of_pow_le hPi)).trans <| Eq.symm <| + (Ring.DimensionLeOne.prime_le_prime_iff_eq (prime j).ne_zero).mp (hPp.le_of_pow_le hPj) open scoped Classical diff --git a/Mathlib/RingTheory/Ideal/Operations.lean b/Mathlib/RingTheory/Ideal/Operations.lean index 35b79eb232f6d..952cf0b1f19a1 100644 --- a/Mathlib/RingTheory/Ideal/Operations.lean +++ b/Mathlib/RingTheory/Ideal/Operations.lean @@ -980,10 +980,37 @@ theorem IsPrime.multiset_prod_mem_iff_exists_mem {I : Ideal R} (hI : I.IsPrime) s.prod ∈ I ↔ ∃ p ∈ s, p ∈ I := by simpa [span_singleton_le_iff_mem] using (hI.multiset_prod_map_le (span {·})) +theorem IsPrime.pow_le_iff {I P : Ideal R} [hP : P.IsPrime] {n : ℕ} (hn : n ≠ 0) : + I ^ n ≤ P ↔ I ≤ P := by + have h : (Multiset.replicate n I).prod ≤ P ↔ _ := hP.multiset_prod_le + simp_rw [Multiset.prod_replicate, Multiset.mem_replicate, ne_eq, hn, not_false_eq_true, + true_and, exists_eq_left] at h + exact h + +@[deprecated (since := "2024-10-06")] alias pow_le_prime_iff := IsPrime.pow_le_iff + +theorem IsPrime.le_of_pow_le {I P : Ideal R} [hP : P.IsPrime] {n : ℕ} (h : I ^ n ≤ P) : + I ≤ P := by + by_cases hn : n = 0 + · rw [hn, pow_zero, one_eq_top] at h + exact fun ⦃_⦄ _ ↦ h Submodule.mem_top + · exact (pow_le_iff hn).mp h + +@[deprecated (since := "2024-10-06")] alias le_of_pow_le_prime := IsPrime.le_of_pow_le + theorem IsPrime.prod_le {s : Finset ι} {f : ι → Ideal R} {P : Ideal R} (hp : IsPrime P) : s.prod f ≤ P ↔ ∃ i ∈ s, f i ≤ P := hp.multiset_prod_map_le f +@[deprecated (since := "2024-10-06")] alias prod_le_prime := IsPrime.prod_le + +/-- The product of a finite number of elements in the commutative semiring `R` lies in the + prime ideal `p` if and only if at least one of those elements is in `p`. -/ +theorem IsPrime.prod_mem_iff {s : Finset ι} {x : ι → R} {p : Ideal R} [hp : p.IsPrime] : + ∏ i in s, x i ∈ p ↔ ∃ i ∈ s, x i ∈ p := by + simp_rw [← span_singleton_le_iff_mem, ← prod_span_singleton] + exact hp.prod_le + theorem IsPrime.prod_mem_iff_exists_mem {I : Ideal R} (hI : I.IsPrime) (s : Finset R) : s.prod (fun x ↦ x) ∈ I ↔ ∃ p ∈ s, p ∈ I := by rw [Finset.prod_eq_multiset_prod, Multiset.map_id'] From e1fa7e2636be2bab66de56099160fdeda4fb5395 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 7 Oct 2024 17:59:20 +0000 Subject: [PATCH 34/43] feat: multiplying an indicator by a constant (#17466) From GibbsMeasure --- Mathlib/Algebra/GroupWithZero/Indicator.lean | 6 ++++++ 1 file changed, 6 insertions(+) diff --git a/Mathlib/Algebra/GroupWithZero/Indicator.lean b/Mathlib/Algebra/GroupWithZero/Indicator.lean index 2b8dcaded4211..7ad962522c252 100644 --- a/Mathlib/Algebra/GroupWithZero/Indicator.lean +++ b/Mathlib/Algebra/GroupWithZero/Indicator.lean @@ -40,6 +40,12 @@ lemma indicator_mul_right (s : Set ι) (f g : ι → M₀) : · rfl · rw [mul_zero] +lemma indicator_mul_const (s : Set ι) (f : ι → M₀) (a : M₀) (i : ι) : + s.indicator (f · * a) i = s.indicator f i * a := by rw [indicator_mul_left] + +lemma indicator_const_mul (s : Set ι) (f : ι → M₀) (a : M₀) (i : ι) : + s.indicator (a * f ·) i = a * s.indicator f i := by rw [indicator_mul_right] + lemma inter_indicator_mul (f g : ι → M₀) (i : ι) : (s ∩ t).indicator (fun j ↦ f j * g j) i = s.indicator f i * t.indicator g i := by rw [← Set.indicator_indicator] From 847ce656a4e7239126dc826d112e3f428db75053 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 7 Oct 2024 17:59:22 +0000 Subject: [PATCH 35/43] feat: supremum of directed indicators (#17467) ... is indicator of the suprema From GibbsMeasure --- Mathlib/Algebra/Order/Group/Indicator.lean | 32 +++++++++++++++++++--- 1 file changed, 28 insertions(+), 4 deletions(-) diff --git a/Mathlib/Algebra/Order/Group/Indicator.lean b/Mathlib/Algebra/Order/Group/Indicator.lean index 6ae4c61cca65e..a8e229c800374 100644 --- a/Mathlib/Algebra/Order/Group/Indicator.lean +++ b/Mathlib/Algebra/Order/Group/Indicator.lean @@ -106,18 +106,25 @@ lemma mulIndicator_le_mulIndicator' (h : a ∈ s → f a ≤ g a) : mulIndicator s f a ≤ mulIndicator s g a := mulIndicator_rel_mulIndicator le_rfl h -@[to_additive] +@[to_additive (attr := mono, gcongr)] lemma mulIndicator_le_mulIndicator (h : f a ≤ g a) : mulIndicator s f a ≤ mulIndicator s g a := mulIndicator_rel_mulIndicator le_rfl fun _ ↦ h -attribute [mono] mulIndicator_le_mulIndicator indicator_le_indicator +@[to_additive (attr := gcongr)] +lemma mulIndicator_mono (h : f ≤ g) : s.mulIndicator f ≤ s.mulIndicator g := + fun _ ↦ mulIndicator_le_mulIndicator (h _) @[to_additive] -lemma mulIndicator_le_mulIndicator_of_subset (h : s ⊆ t) (hf : ∀ a, 1 ≤ f a) (a : α) : +lemma mulIndicator_le_mulIndicator_apply_of_subset (h : s ⊆ t) (hf : 1 ≤ f a) : mulIndicator s f a ≤ mulIndicator t f a := mulIndicator_apply_le' (fun ha ↦ le_mulIndicator_apply (fun _ ↦ le_rfl) fun hat ↦ (hat <| h ha).elim) fun _ ↦ - one_le_mulIndicator_apply fun _ ↦ hf _ + one_le_mulIndicator_apply fun _ ↦ hf + +@[to_additive] +lemma mulIndicator_le_mulIndicator_of_subset (h : s ⊆ t) (hf : 1 ≤ f) : + mulIndicator s f ≤ mulIndicator t f := + fun _ ↦ mulIndicator_le_mulIndicator_apply_of_subset h (hf _) @[to_additive] lemma mulIndicator_le_self' (hf : ∀ x ∉ s, 1 ≤ f x) : mulIndicator s f ≤ f := @@ -174,6 +181,23 @@ lemma mulIndicator_iInter_apply (h1 : (⊥ : M) = 1) (s : ι → Set α) (f : α refine le_antisymm (by simp only [← h1, le_iInf_iff, bot_le, forall_const]) ?_ simpa [mulIndicator_of_not_mem hj] using (iInf_le (fun i ↦ (s i).mulIndicator f) j) x +@[to_additive] +lemma iSup_mulIndicator {ι : Type*} [Preorder ι] [IsDirected ι (· ≤ ·)] {f : ι → α → M} + {s : ι → Set α} (h1 : (⊥ : M) = 1) (hf : Monotone f) (hs : Monotone s) : + ⨆ i, (s i).mulIndicator (f i) = (⋃ i, s i).mulIndicator (⨆ i, f i) := by + simp only [le_antisymm_iff, iSup_le_iff] + refine ⟨fun i ↦ (mulIndicator_mono (le_iSup _ _)).trans (mulIndicator_le_mulIndicator_of_subset + (subset_iUnion _ _) (fun _ ↦ by simp [← h1])), fun a ↦ ?_⟩ + by_cases ha : a ∈ ⋃ i, s i + · obtain ⟨i, hi⟩ : ∃ i, a ∈ s i := by simpa using ha + rw [mulIndicator_of_mem ha, iSup_apply, iSup_apply] + refine iSup_le fun j ↦ ?_ + obtain ⟨k, hik, hjk⟩ := exists_ge_ge i j + refine le_iSup_of_le k <| (hf hjk _).trans_eq ?_ + rw [mulIndicator_of_mem (hs hik hi)] + · rw [mulIndicator_of_not_mem ha, ← h1] + exact bot_le + end CompleteLattice section CanonicallyOrderedCommMonoid From 6fc3b64cb05f32b8f0394adada6d848c9f3e5109 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 7 Oct 2024 17:59:23 +0000 Subject: [PATCH 36/43] chore(Topology): rename some `prod_map`/`prod` to `prodMap` (#17472) Also - replace some proofs by `fun_prop` - use `Prod.map` instead of an inline lambda in some lemmas - fix `to_additive` name for `ContinuousAddMonoidHom.prod` and `ContinuousAddMonoidHom.prodMap` (were `sum` and `sum_map`). --- Counterexamples/SorgenfreyLine.lean | 2 +- Mathlib/Analysis/Calculus/FDeriv/Comp.lean | 2 +- Mathlib/Analysis/Calculus/FDeriv/Equiv.lean | 4 +- Mathlib/Analysis/InnerProductSpace/Basic.lean | 6 +- .../Analysis/Normed/Module/Completion.lean | 5 +- .../NormedSpace/Multilinear/Basic.lean | 4 +- .../SpecialFunctions/Pow/Continuity.lean | 9 +-- .../DynamicalEntourage.lean | 2 +- .../Manifold/SmoothManifoldWithCorners.lean | 4 +- .../Constructions/Polish/Basic.lean | 2 +- .../Topology/Algebra/ContinuousMonoidHom.lean | 22 ++++--- Mathlib/Topology/Algebra/Group/Basic.lean | 4 +- Mathlib/Topology/Algebra/Module/Basic.lean | 4 +- Mathlib/Topology/Algebra/Monoid.lean | 2 +- Mathlib/Topology/Algebra/MulAction.lean | 4 +- Mathlib/Topology/Algebra/Order/Floor.lean | 2 +- Mathlib/Topology/Algebra/ProperAction.lean | 4 +- Mathlib/Topology/Algebra/ProperConstSMul.lean | 2 +- Mathlib/Topology/Algebra/UniformGroup.lean | 6 +- Mathlib/Topology/Algebra/UniformRing.lean | 4 +- Mathlib/Topology/CompactOpen.lean | 4 +- Mathlib/Topology/Connected/PathConnected.lean | 10 ++-- Mathlib/Topology/Constructions.lean | 60 ++++++++++++------- Mathlib/Topology/ContinuousMap/Algebra.lean | 6 +- Mathlib/Topology/ContinuousMap/Basic.lean | 2 - Mathlib/Topology/DenseEmbedding.lean | 21 ++++--- .../Topology/FiberBundle/Constructions.lean | 10 ++-- Mathlib/Topology/Homeomorph.lean | 12 ++-- Mathlib/Topology/Homotopy/HomotopyGroup.lean | 2 +- Mathlib/Topology/Inseparable.lean | 11 +--- Mathlib/Topology/Instances/ENNReal.lean | 2 +- Mathlib/Topology/Instances/EReal.lean | 4 +- Mathlib/Topology/Instances/Rat.lean | 2 +- Mathlib/Topology/Maps/Proper/Basic.lean | 8 ++- Mathlib/Topology/Order/OrderClosed.lean | 2 +- Mathlib/Topology/Separation.lean | 2 +- .../UniformSpace/AbstractCompletion.lean | 2 +- Mathlib/Topology/UniformSpace/Basic.lean | 14 +++-- Mathlib/Topology/UniformSpace/Compact.lean | 2 +- Mathlib/Topology/UniformSpace/Completion.lean | 4 +- test/Continuity.lean | 3 +- 41 files changed, 147 insertions(+), 130 deletions(-) diff --git a/Counterexamples/SorgenfreyLine.lean b/Counterexamples/SorgenfreyLine.lean index 746700b0769de..dc20c635dcc46 100644 --- a/Counterexamples/SorgenfreyLine.lean +++ b/Counterexamples/SorgenfreyLine.lean @@ -141,7 +141,7 @@ theorem continuous_toReal : Continuous toReal := exact inf_le_left instance : OrderClosedTopology ℝₗ := - ⟨isClosed_le_prod.preimage (continuous_toReal.prod_map continuous_toReal)⟩ + ⟨isClosed_le_prod.preimage (continuous_toReal.prodMap continuous_toReal)⟩ instance : ContinuousAdd ℝₗ := by refine ⟨continuous_iff_continuousAt.2 ?_⟩ diff --git a/Mathlib/Analysis/Calculus/FDeriv/Comp.lean b/Mathlib/Analysis/Calculus/FDeriv/Comp.lean index beaf6fad6b9fe..410d6d1cf3e69 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Comp.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Comp.lean @@ -172,7 +172,7 @@ theorem Differentiable.comp_differentiableOn {g : F → G} (hg : Differentiable protected theorem HasStrictFDerivAt.comp {g : F → G} {g' : F →L[𝕜] G} (hg : HasStrictFDerivAt g g' (f x)) (hf : HasStrictFDerivAt f f' x) : HasStrictFDerivAt (fun x => g (f x)) (g'.comp f') x := - ((hg.comp_tendsto (hf.continuousAt.prod_map' hf.continuousAt)).trans_isBigO + ((hg.comp_tendsto (hf.continuousAt.prodMap' hf.continuousAt)).trans_isBigO hf.isBigO_sub).triangle <| by simpa only [g'.map_sub, f'.coe_comp'] using (g'.isBigO_comp _ _).trans_isLittleO hf diff --git a/Mathlib/Analysis/Calculus/FDeriv/Equiv.lean b/Mathlib/Analysis/Calculus/FDeriv/Equiv.lean index 1c6738f520fcc..ac47ac267e108 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Equiv.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Equiv.lean @@ -339,7 +339,7 @@ inverse function. -/ theorem HasStrictFDerivAt.of_local_left_inverse {f : E → F} {f' : E ≃L[𝕜] F} {g : F → E} {a : F} (hg : ContinuousAt g a) (hf : HasStrictFDerivAt f (f' : E →L[𝕜] F) (g a)) (hfg : ∀ᶠ y in 𝓝 a, f (g y) = y) : HasStrictFDerivAt g (f'.symm : F →L[𝕜] E) a := by - replace hg := hg.prod_map' hg + replace hg := hg.prodMap' hg replace hfg := hfg.prod_mk_nhds hfg have : (fun p : F × F => g p.1 - g p.2 - f'.symm (p.1 - p.2)) =O[𝓝 (a, a)] fun p : F × F => @@ -355,7 +355,7 @@ theorem HasStrictFDerivAt.of_local_left_inverse {f : E → F} {f' : E ≃L[𝕜] · refine (hf.isBigO_sub_rev.comp_tendsto hg).congr' (Eventually.of_forall fun _ => rfl) (hfg.mono ?_) rintro p ⟨hp1, hp2⟩ - simp only [(· ∘ ·), hp1, hp2] + simp only [(· ∘ ·), hp1, hp2, Prod.map] /-- 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`, then `g` has the derivative `f'⁻¹` at `a`. diff --git a/Mathlib/Analysis/InnerProductSpace/Basic.lean b/Mathlib/Analysis/InnerProductSpace/Basic.lean index b077cc931cc99..d4672630df1ed 100644 --- a/Mathlib/Analysis/InnerProductSpace/Basic.lean +++ b/Mathlib/Analysis/InnerProductSpace/Basic.lean @@ -2384,11 +2384,11 @@ open UniformSpace Function instance toInner {𝕜' E' : Type*} [TopologicalSpace 𝕜'] [UniformSpace E'] [Inner 𝕜' E'] : Inner 𝕜' (Completion E') where - inner := curry <| (isDenseInducing_coe.prod isDenseInducing_coe).extend (uncurry inner) + inner := curry <| (isDenseInducing_coe.prodMap isDenseInducing_coe).extend (uncurry inner) @[simp] theorem inner_coe (a b : E) : inner (a : Completion E) (b : Completion E) = (inner a b : 𝕜) := - (isDenseInducing_coe.prod isDenseInducing_coe).extend_eq + (isDenseInducing_coe.prodMap isDenseInducing_coe).extend_eq (continuous_inner : Continuous (uncurry inner : E × E → 𝕜)) (a, b) protected theorem continuous_inner : @@ -2401,7 +2401,7 @@ protected theorem continuous_inner : rw [Completion.toInner, inner, uncurry_curry _] change Continuous - (((isDenseInducing_toCompl E).prod (isDenseInducing_toCompl E)).extend fun p : E × E => + (((isDenseInducing_toCompl E).prodMap (isDenseInducing_toCompl E)).extend fun p : E × E => inner' p.1 p.2) exact (isDenseInducing_toCompl E).extend_Z_bilin (isDenseInducing_toCompl E) this diff --git a/Mathlib/Analysis/Normed/Module/Completion.lean b/Mathlib/Analysis/Normed/Module/Completion.lean index c4377b5b2a68a..f0a8fdeed667d 100644 --- a/Mathlib/Analysis/Normed/Module/Completion.lean +++ b/Mathlib/Analysis/Normed/Module/Completion.lean @@ -74,10 +74,7 @@ instance [SeminormedRing A] : NormedRing (Completion A) where norm_mul x y := by induction x, y using induction_on₂ with | hp => - exact - isClosed_le (Continuous.comp continuous_norm continuous_mul) - (Continuous.comp _root_.continuous_mul - (Continuous.prod_map continuous_norm continuous_norm)) + apply isClosed_le <;> fun_prop | ih x y => simp only [← coe_mul, norm_coe] exact norm_mul_le x y diff --git a/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean b/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean index 898e770696c72..4b87da345fe31 100644 --- a/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean +++ b/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean @@ -87,7 +87,7 @@ theorem ContinuousMultilinearMap.continuous_eval : let _ := TopologicalAddGroup.toUniformSpace F have := comm_topologicalAddGroup_is_uniform (G := F) refine (UniformOnFun.continuousOn_eval₂ fun m ↦ ?_).comp_continuous - (embedding_toUniformOnFun.continuous.prod_map continuous_id) fun (f, x) ↦ f.cont.continuousAt + (embedding_toUniformOnFun.continuous.prodMap continuous_id) fun (f, x) ↦ f.cont.continuousAt exact ⟨ball m 1, NormedSpace.isVonNBounded_of_isBounded _ isBounded_ball, ball_mem_nhds _ one_pos⟩ @@ -98,7 +98,7 @@ variable {G : Type*} [AddCommGroup G] [TopologicalSpace G] [Module 𝕜 G] [Cont lemma continuous_uncurry_of_multilinear : Continuous (fun (p : G × (Π i, E i)) ↦ f p.1 p.2) := - ContinuousMultilinearMap.continuous_eval.comp <| .prod_map (map_continuous f) continuous_id + ContinuousMultilinearMap.continuous_eval.comp <| .prodMap (map_continuous f) continuous_id lemma continuousOn_uncurry_of_multilinear {s} : ContinuousOn (fun (p : G × (Π i, E i)) ↦ f p.1 p.2) s := diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Continuity.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Continuity.lean index d3ff288d66b92..841d4fc4de757 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Continuity.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Continuity.lean @@ -332,8 +332,7 @@ theorem continuousAt_ofReal_cpow (x : ℝ) (y : ℂ) (h : 0 < y.re ∨ x ≠ 0) ContinuousAt (fun p => (p.1 : ℂ) ^ p.2 : ℝ × ℂ → ℂ) (x, y) := by rcases lt_trichotomy (0 : ℝ) x with (hx | rfl | hx) · -- x > 0 : easy case - have : ContinuousAt (fun p => ⟨↑p.1, p.2⟩ : ℝ × ℂ → ℂ × ℂ) (x, y) := - continuous_ofReal.continuousAt.prod_map continuousAt_id + have : ContinuousAt (fun p => ⟨↑p.1, p.2⟩ : ℝ × ℂ → ℂ × ℂ) (x, y) := by fun_prop refine (continuousAt_cpow (Or.inl ?_)).comp this rwa [ofReal_re] · -- x = 0 : reduce to continuousAt_cpow_zero_of_re_pos @@ -341,15 +340,13 @@ theorem continuousAt_ofReal_cpow (x : ℝ) (y : ℂ) (h : 0 < y.re ∨ x ≠ 0) rw [ofReal_zero] apply continuousAt_cpow_zero_of_re_pos tauto - have B : ContinuousAt (fun p => ⟨↑p.1, p.2⟩ : ℝ × ℂ → ℂ × ℂ) ⟨0, y⟩ := - continuous_ofReal.continuousAt.prod_map continuousAt_id + have B : ContinuousAt (fun p => ⟨↑p.1, p.2⟩ : ℝ × ℂ → ℂ × ℂ) ⟨0, y⟩ := by fun_prop exact A.comp_of_eq B rfl · -- x < 0 : difficult case suffices ContinuousAt (fun p => (-(p.1 : ℂ)) ^ p.2 * exp (π * I * p.2) : ℝ × ℂ → ℂ) (x, y) by refine this.congr (eventually_of_mem (prod_mem_nhds (Iio_mem_nhds hx) univ_mem) ?_) exact fun p hp => (ofReal_cpow_of_nonpos (le_of_lt hp.1) p.2).symm - have A : ContinuousAt (fun p => ⟨-↑p.1, p.2⟩ : ℝ × ℂ → ℂ × ℂ) (x, y) := - ContinuousAt.prod_map continuous_ofReal.continuousAt.neg continuousAt_id + have A : ContinuousAt (fun p => ⟨-↑p.1, p.2⟩ : ℝ × ℂ → ℂ × ℂ) (x, y) := by fun_prop apply ContinuousAt.mul · refine (continuousAt_cpow (Or.inl ?_)).comp A rwa [neg_re, ofReal_re, neg_pos] diff --git a/Mathlib/Dynamics/TopologicalEntropy/DynamicalEntourage.lean b/Mathlib/Dynamics/TopologicalEntropy/DynamicalEntourage.lean index ad6c52b6be339..08b6c75d0d4c0 100644 --- a/Mathlib/Dynamics/TopologicalEntropy/DynamicalEntourage.lean +++ b/Mathlib/Dynamics/TopologicalEntropy/DynamicalEntourage.lean @@ -88,7 +88,7 @@ lemma _root_.isOpen.dynEntourage [TopologicalSpace X] {T : X → X} (T_cont : Co IsOpen (dynEntourage T U n) := by rw [dynEntourage_eq_inter_Ico T U n] refine isOpen_iInter_of_finite fun k ↦ ?_ - exact continuous_def.1 ((T_cont.prod_map T_cont).iterate k) U U_open + exact U_open.preimage ((T_cont.prodMap T_cont).iterate k) lemma dynEntourage_monotone (T : X → X) (n : ℕ) : Monotone (fun U : Set (X × X) ↦ dynEntourage T U n) := diff --git a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean index 9f6ab2d09bb29..1b1e09a291612 100644 --- a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean +++ b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean @@ -382,8 +382,8 @@ def ModelWithCorners.prod {𝕜 : Type u} [NontriviallyNormedField 𝕜] {E : Ty source := { x | x.1 ∈ I.source ∧ x.2 ∈ I'.source } source_eq := by simp only [setOf_true, mfld_simps] uniqueDiffOn' := I.uniqueDiffOn'.prod I'.uniqueDiffOn' - continuous_toFun := I.continuous_toFun.prod_map I'.continuous_toFun - continuous_invFun := I.continuous_invFun.prod_map I'.continuous_invFun } + continuous_toFun := I.continuous_toFun.prodMap I'.continuous_toFun + continuous_invFun := I.continuous_invFun.prodMap I'.continuous_invFun } /-- Given a finite family of `ModelWithCorners` `I i` on `(E i, H i)`, we define the model with corners `pi I` on `(Π i, E i, ModelPi H)`. See note [Manifold type tags] for explanation about diff --git a/Mathlib/MeasureTheory/Constructions/Polish/Basic.lean b/Mathlib/MeasureTheory/Constructions/Polish/Basic.lean index f1677d8c92db5..89db09097d3cb 100644 --- a/Mathlib/MeasureTheory/Constructions/Polish/Basic.lean +++ b/Mathlib/MeasureTheory/Constructions/Polish/Basic.lean @@ -343,7 +343,7 @@ protected lemma AnalyticSet.preimage {X Y : Type*} [TopologicalSpace X] [Topolog [PolishSpace X] [T2Space Y] {s : Set Y} (hs : AnalyticSet s) {f : X → Y} (hf : Continuous f) : AnalyticSet (f ⁻¹' s) := by rcases analyticSet_iff_exists_polishSpace_range.1 hs with ⟨Z, _, _, g, hg, rfl⟩ - have : IsClosed {x : X × Z | f x.1 = g x.2} := isClosed_diagonal.preimage (hf.prod_map hg) + have : IsClosed {x : X × Z | f x.1 = g x.2} := isClosed_eq hf.fst' hg.snd' convert this.analyticSet.image_of_continuous continuous_fst ext x simp [eq_comm] diff --git a/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean b/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean index 00b7d7e857345..a753c9397ac13 100644 --- a/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean +++ b/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean @@ -131,16 +131,24 @@ def comp (g : ContinuousMonoidHom B C) (f : ContinuousMonoidHom A B) : Continuou mk' (g.toMonoidHom.comp f.toMonoidHom) (g.continuous_toFun.comp f.continuous_toFun) /-- Product of two continuous homomorphisms on the same space. -/ -@[to_additive (attr := simps!) "Product of two continuous homomorphisms on the same space."] +@[to_additive (attr := simps!) prod "Product of two continuous homomorphisms on the same space."] def prod (f : ContinuousMonoidHom A B) (g : ContinuousMonoidHom A C) : ContinuousMonoidHom A (B × C) := mk' (f.toMonoidHom.prod g.toMonoidHom) (f.continuous_toFun.prod_mk g.continuous_toFun) /-- Product of two continuous homomorphisms on different spaces. -/ -@[to_additive (attr := simps!) "Product of two continuous homomorphisms on different spaces."] -def prod_map (f : ContinuousMonoidHom A C) (g : ContinuousMonoidHom B D) : +@[to_additive (attr := simps!) prodMap + "Product of two continuous homomorphisms on different spaces."] +def prodMap (f : ContinuousMonoidHom A C) (g : ContinuousMonoidHom B D) : ContinuousMonoidHom (A × B) (C × D) := - mk' (f.toMonoidHom.prodMap g.toMonoidHom) (f.continuous_toFun.prod_map g.continuous_toFun) + mk' (f.toMonoidHom.prodMap g.toMonoidHom) (f.continuous_toFun.prodMap g.continuous_toFun) + +@[deprecated (since := "2024-10-05")] alias prod_map := prodMap +@[deprecated (since := "2024-10-05")] +alias _root_.ContinuousAddMonoidHom.sum_map := ContinuousAddMonoidHom.prodMap + +set_option linter.existingAttributeWarning false in +attribute [to_additive existing] prod_map variable (A B C D E) @@ -208,7 +216,7 @@ variable {A B C D E} @[to_additive (attr := simps!) "Coproduct of two continuous homomorphisms to the same space."] def coprod (f : ContinuousMonoidHom A E) (g : ContinuousMonoidHom B E) : ContinuousMonoidHom (A × B) E := - (mul E).comp (f.prod_map g) + (mul E).comp (f.prodMap g) @[to_additive] instance : CommGroup (ContinuousMonoidHom A E) where @@ -265,7 +273,7 @@ instance [T2Space B] : T2Space (ContinuousMonoidHom A B) := instance : TopologicalGroup (ContinuousMonoidHom A E) := let hi := inducing_toContinuousMap A E let hc := hi.continuous - { continuous_mul := hi.continuous_iff.mpr (continuous_mul.comp (Continuous.prod_map hc hc)) + { continuous_mul := hi.continuous_iff.mpr (continuous_mul.comp (Continuous.prodMap hc hc)) continuous_inv := hi.continuous_iff.mpr (continuous_inv.comp hc) } @[to_additive] @@ -280,7 +288,7 @@ theorem continuous_comp [LocallyCompactSpace B] : Continuous fun f : ContinuousMonoidHom A B × ContinuousMonoidHom B C => f.2.comp f.1 := (inducing_toContinuousMap A C).continuous_iff.2 <| ContinuousMap.continuous_comp'.comp - ((inducing_toContinuousMap A B).prod_map (inducing_toContinuousMap B C)).continuous + ((inducing_toContinuousMap A B).prodMap (inducing_toContinuousMap B C)).continuous @[to_additive] theorem continuous_comp_left (f : ContinuousMonoidHom A B) : diff --git a/Mathlib/Topology/Algebra/Group/Basic.lean b/Mathlib/Topology/Algebra/Group/Basic.lean index 99386fcf912e5..49c268a901a0c 100644 --- a/Mathlib/Topology/Algebra/Group/Basic.lean +++ b/Mathlib/Topology/Algebra/Group/Basic.lean @@ -536,7 +536,7 @@ end OrderedCommGroup @[to_additive] instance [TopologicalSpace H] [Group H] [TopologicalGroup H] : TopologicalGroup (G × H) where - continuous_inv := continuous_inv.prod_map continuous_inv + continuous_inv := continuous_inv.prodMap continuous_inv @[to_additive] instance Pi.topologicalGroup {C : β → Type*} [∀ b, TopologicalSpace (C b)] [∀ b, Group (C b)] @@ -1426,7 +1426,7 @@ theorem Subgroup.properlyDiscontinuousSMul_opposite_of_tendsto_cofinite (S : Sub (hS : Tendsto S.subtype cofinite (cocompact G)) : ProperlyDiscontinuousSMul S.op G := { finite_disjoint_inter_image := by intro K L hK hL - have : Continuous fun p : G × G => (p.1⁻¹, p.2) := continuous_inv.prod_map continuous_id + have : Continuous fun p : G × G => (p.1⁻¹, p.2) := continuous_inv.prodMap continuous_id have H : Set.Finite _ := hS ((hK.prod hL).image (continuous_mul.comp this)).compl_mem_cocompact simp only [preimage_compl, compl_compl, coeSubtype, comp_apply] at H diff --git a/Mathlib/Topology/Algebra/Module/Basic.lean b/Mathlib/Topology/Algebra/Module/Basic.lean index 76d6e898d399b..db5489036d28f 100644 --- a/Mathlib/Topology/Algebra/Module/Basic.lean +++ b/Mathlib/Topology/Algebra/Module/Basic.lean @@ -1817,8 +1817,8 @@ theorem trans_toLinearEquiv (e₁ : M₁ ≃SL[σ₁₂] M₂) (e₂ : M₂ ≃S def prod [Module R₁ M₂] [Module R₁ M₃] [Module R₁ M₄] (e : M₁ ≃L[R₁] M₂) (e' : M₃ ≃L[R₁] M₄) : (M₁ × M₃) ≃L[R₁] M₂ × M₄ := { e.toLinearEquiv.prod e'.toLinearEquiv with - continuous_toFun := e.continuous_toFun.prod_map e'.continuous_toFun - continuous_invFun := e.continuous_invFun.prod_map e'.continuous_invFun } + continuous_toFun := e.continuous_toFun.prodMap e'.continuous_toFun + continuous_invFun := e.continuous_invFun.prodMap e'.continuous_invFun } @[simp, norm_cast] theorem prod_apply [Module R₁ M₂] [Module R₁ M₃] [Module R₁ M₄] (e : M₁ ≃L[R₁] M₂) diff --git a/Mathlib/Topology/Algebra/Monoid.lean b/Mathlib/Topology/Algebra/Monoid.lean index a8001b461554b..77b696c9d052c 100644 --- a/Mathlib/Topology/Algebra/Monoid.lean +++ b/Mathlib/Topology/Algebra/Monoid.lean @@ -75,7 +75,7 @@ instance ContinuousMul.to_continuousSMul : ContinuousSMul M M := instance ContinuousMul.to_continuousSMul_op : ContinuousSMul Mᵐᵒᵖ M := ⟨show Continuous ((fun p : M × M => p.1 * p.2) ∘ Prod.swap ∘ Prod.map MulOpposite.unop id) from continuous_mul.comp <| - continuous_swap.comp <| Continuous.prod_map MulOpposite.continuous_unop continuous_id⟩ + continuous_swap.comp <| Continuous.prodMap MulOpposite.continuous_unop continuous_id⟩ @[to_additive] theorem ContinuousMul.induced {α : Type*} {β : Type*} {F : Type*} [FunLike F α β] [MulOneClass α] diff --git a/Mathlib/Topology/Algebra/MulAction.lean b/Mathlib/Topology/Algebra/MulAction.lean index b17c9a9e4bb1a..91dc843aa0507 100644 --- a/Mathlib/Topology/Algebra/MulAction.lean +++ b/Mathlib/Topology/Algebra/MulAction.lean @@ -132,13 +132,13 @@ action is."] instance ContinuousSMul.op [SMul Mᵐᵒᵖ X] [IsCentralScalar M X] : ContinuousSMul Mᵐᵒᵖ X := ⟨by suffices Continuous fun p : M × X => MulOpposite.op p.fst • p.snd from - this.comp (MulOpposite.continuous_unop.prod_map continuous_id) + this.comp (MulOpposite.continuous_unop.prodMap continuous_id) simpa only [op_smul_eq_smul] using (continuous_smul : Continuous fun p : M × X => _)⟩ @[to_additive] instance MulOpposite.continuousSMul : ContinuousSMul M Xᵐᵒᵖ := ⟨MulOpposite.continuous_op.comp <| - continuous_smul.comp <| continuous_id.prod_map MulOpposite.continuous_unop⟩ + continuous_smul.comp <| continuous_id.prodMap MulOpposite.continuous_unop⟩ @[to_additive] protected theorem Specializes.smul {a b : M} {x y : X} (h₁ : a ⤳ b) (h₂ : x ⤳ y) : diff --git a/Mathlib/Topology/Algebra/Order/Floor.lean b/Mathlib/Topology/Algebra/Order/Floor.lean index 390b871f2f357..bde8ae91b5715 100644 --- a/Mathlib/Topology/Algebra/Order/Floor.lean +++ b/Mathlib/Topology/Algebra/Order/Floor.lean @@ -206,7 +206,7 @@ theorem ContinuousOn.comp_fract' {f : β → α → γ} (h : ContinuousOn (uncur (tendsto_id.prod_map (tendsto_fract_right _))).mono_right (le_of_eq ?_) <;> simp [nhdsWithin_prod_eq, nhdsWithin_univ] · replace ht : t ≠ ⌊t⌋ := fun ht' => ht ⟨_, ht'⟩ - refine (h.continuousAt ?_).comp (continuousAt_id.prod_map (continuousAt_fract ht)) + refine (h.continuousAt ?_).comp (continuousAt_id.prodMap (continuousAt_fract ht)) exact prod_mem_nhds univ_mem (Icc_mem_nhds (fract_pos.2 ht) (fract_lt_one _)) theorem ContinuousOn.comp_fract {s : β → α} {f : β → α → γ} diff --git a/Mathlib/Topology/Algebra/ProperAction.lean b/Mathlib/Topology/Algebra/ProperAction.lean index b937385d62d95..d37212b162f3c 100644 --- a/Mathlib/Topology/Algebra/ProperAction.lean +++ b/Mathlib/Topology/Algebra/ProperAction.lean @@ -173,7 +173,7 @@ theorem properSMul_of_closedEmbedding {H : Type*} [Group H] [MulAction H X] [Top (f_compat : ∀ (h : H) (x : X), f h • x = h • x) : ProperSMul H X where isProperMap_smul_pair := by have := isProperMap_of_closedEmbedding f_clemb - have h : IsProperMap (Prod.map f (fun x : X ↦ x)) := IsProperMap.prod_map this isProperMap_id + have h : IsProperMap (Prod.map f (fun x : X ↦ x)) := this.prodMap isProperMap_id have : (fun hx : H × X ↦ (hx.1 • hx.2, hx.2)) = (fun hx ↦ (f hx.1 • hx.2, hx.2)) := by simp [f_compat] rw [this] @@ -245,7 +245,7 @@ theorem properlyDiscontinuousSMul_iff_properSMul [T2Space X] [DiscreteTopology G apply IsCompact.finite_of_discrete -- Now set `h : (g, x) ↦ (g⁻¹ • x, x)`, because `f` is proper by hypothesis, so is `h`. have : IsProperMap (fun gx : G × X ↦ (gx.1⁻¹ • gx.2, gx.2)) := - (IsProperMap.prod_map (Homeomorph.isProperMap (Homeomorph.inv G)) isProperMap_id).comp <| + (IsProperMap.prodMap (Homeomorph.isProperMap (Homeomorph.inv G)) isProperMap_id).comp <| ProperSMul.isProperMap_smul_pair --But we also have that `{g | Set.Nonempty ((g • ·) '' K ∩ L)} = h ⁻¹ (K × L)`, which -- concludes the proof. diff --git a/Mathlib/Topology/Algebra/ProperConstSMul.lean b/Mathlib/Topology/Algebra/ProperConstSMul.lean index 30fbbb4b83bf7..7e3167890f5ee 100644 --- a/Mathlib/Topology/Algebra/ProperConstSMul.lean +++ b/Mathlib/Topology/Algebra/ProperConstSMul.lean @@ -61,7 +61,7 @@ instance {M X Y : Type*} [SMul M X] [TopologicalSpace X] [ProperConstSMul M X] [SMul M Y] [TopologicalSpace Y] [ProperConstSMul M Y] : ProperConstSMul M (X × Y) := - ⟨fun c ↦ (isProperMap_smul c X).prod_map (isProperMap_smul c Y)⟩ + ⟨fun c ↦ (isProperMap_smul c X).prodMap (isProperMap_smul c Y)⟩ instance {M ι : Type*} {X : ι → Type*} [∀ i, SMul M (X i)] [∀ i, TopologicalSpace (X i)] [∀ i, ProperConstSMul M (X i)] : diff --git a/Mathlib/Topology/Algebra/UniformGroup.lean b/Mathlib/Topology/Algebra/UniformGroup.lean index daf8079c33a30..a8167d4c22d8d 100644 --- a/Mathlib/Topology/Algebra/UniformGroup.lean +++ b/Mathlib/Topology/Algebra/UniformGroup.lean @@ -231,7 +231,7 @@ lemma UniformInducing.uniformGroup {γ : Type*} [Group γ] [UniformSpace γ] [Un UniformGroup β where uniformContinuous_div := by simp_rw [hf.uniformContinuous_iff, Function.comp_def, map_div] - exact uniformContinuous_div.comp (hf.uniformContinuous.prod_map hf.uniformContinuous) + exact uniformContinuous_div.comp (hf.uniformContinuous.prodMap hf.uniformContinuous) @[to_additive] protected theorem UniformGroup.comap {γ : Type*} [Group γ] {u : UniformSpace γ} [UniformGroup γ] @@ -735,14 +735,14 @@ variable [T0Space G] [CompleteSpace G] /-- Bourbaki GT III.6.5 Theorem I: ℤ-bilinear continuous maps from dense images into a complete Hausdorff group extend by continuity. Note: Bourbaki assumes that α and β are also complete Hausdorff, but this is not necessary. -/ -theorem extend_Z_bilin : Continuous (extend (de.prod df) (fun p : β × δ => φ p.1 p.2)) := by +theorem extend_Z_bilin : Continuous (extend (de.prodMap df) (fun p : β × δ => φ p.1 p.2)) := by refine continuous_extend_of_cauchy _ ?_ rintro ⟨x₀, y₀⟩ constructor · apply NeBot.map apply comap_neBot intro U h - rcases mem_closure_iff_nhds.1 ((de.prod df).dense (x₀, y₀)) U h with ⟨x, x_in, ⟨z, z_x⟩⟩ + rcases mem_closure_iff_nhds.1 ((de.prodMap df).dense (x₀, y₀)) U h with ⟨x, x_in, ⟨z, z_x⟩⟩ exists z aesop · suffices map (fun p : (β × δ) × β × δ => (fun p : β × δ => φ p.1 p.2) p.2 - diff --git a/Mathlib/Topology/Algebra/UniformRing.lean b/Mathlib/Topology/Algebra/UniformRing.lean index f4799994dceaf..81be27c4b6eab 100644 --- a/Mathlib/Topology/Algebra/UniformRing.lean +++ b/Mathlib/Topology/Algebra/UniformRing.lean @@ -47,7 +47,7 @@ instance one : One (Completion α) := ⟨(1 : α)⟩ instance mul : Mul (Completion α) := - ⟨curry <| (isDenseInducing_coe.prod isDenseInducing_coe).extend ((↑) ∘ uncurry (· * ·))⟩ + ⟨curry <| (isDenseInducing_coe.prodMap isDenseInducing_coe).extend ((↑) ∘ uncurry (· * ·))⟩ @[norm_cast] theorem coe_one : ((1 : α) : Completion α) = 1 := @@ -59,7 +59,7 @@ variable {α : Type*} [Ring α] [UniformSpace α] [TopologicalRing α] @[norm_cast] theorem coe_mul (a b : α) : ((a * b : α) : Completion α) = a * b := - ((isDenseInducing_coe.prod isDenseInducing_coe).extend_eq + ((isDenseInducing_coe.prodMap isDenseInducing_coe).extend_eq ((continuous_coe α).comp (@continuous_mul α _ _ _)) (a, b)).symm variable [UniformAddGroup α] diff --git a/Mathlib/Topology/CompactOpen.lean b/Mathlib/Topology/CompactOpen.lean index 6462eef116241..0338dd8e2d361 100644 --- a/Mathlib/Topology/CompactOpen.lean +++ b/Mathlib/Topology/CompactOpen.lean @@ -374,7 +374,7 @@ theorem continuous_curry [LocallyCompactSpace (X × Y)] : /-- The uncurried form of a continuous map `X → C(Y, Z)` is a continuous map `X × Y → Z`. -/ theorem continuous_uncurry_of_continuous [LocallyCompactSpace Y] (f : C(X, C(Y, Z))) : Continuous (Function.uncurry fun x y => f x y) := - continuous_eval.comp <| f.continuous.prod_map continuous_id + continuous_eval.comp <| f.continuous.prodMap continuous_id /-- The uncurried form of a continuous map `X → C(Y, Z)` as a continuous map `X × Y → Z` (if `Y` is locally compact). If `X` is also locally compact, then this is a homeomorphism between the two @@ -388,7 +388,7 @@ theorem continuous_uncurry [LocallyCompactSpace X] [LocallyCompactSpace Y] : Continuous (uncurry : C(X, C(Y, Z)) → C(X × Y, Z)) := by apply continuous_of_continuous_uncurry rw [← (Homeomorph.prodAssoc _ _ _).comp_continuous_iff'] - apply continuous_eval.comp (continuous_eval.prod_map continuous_id) + apply continuous_eval.comp (continuous_eval.prodMap continuous_id) /-- The family of constant maps: `Y → C(X, Y)` as a continuous map. -/ def const' : C(Y, C(X, Y)) := diff --git a/Mathlib/Topology/Connected/PathConnected.lean b/Mathlib/Topology/Connected/PathConnected.lean index 2d89d2ed72042..644fcb694b3a8 100644 --- a/Mathlib/Topology/Connected/PathConnected.lean +++ b/Mathlib/Topology/Connected/PathConnected.lean @@ -200,7 +200,7 @@ instance topologicalSpace : TopologicalSpace (Path x y) := theorem continuous_eval : Continuous fun p : Path x y × I => p.1 p.2 := ContinuousMap.continuous_eval.comp <| - (continuous_induced_dom (α := Path x y)).prod_map continuous_id + (continuous_induced_dom (α := Path x y)).prodMap continuous_id @[continuity] theorem _root_.Continuous.path_eval {Y} [TopologicalSpace Y] {f : Y → Path x y} {g : Y → I} @@ -423,7 +423,7 @@ theorem cast_coe (γ : Path x y) {x' y'} (hx : x' = x) (hy : y' = y) : (γ.cast theorem symm_continuous_family {ι : Type*} [TopologicalSpace ι] {a b : ι → X} (γ : ∀ t : ι, Path (a t) (b t)) (h : Continuous ↿γ) : Continuous ↿fun t => (γ t).symm := - h.comp (continuous_id.prod_map continuous_symm) + h.comp (continuous_id.prodMap continuous_symm) @[continuity] theorem continuous_symm : Continuous (symm : Path x y → Path y x) := @@ -433,7 +433,7 @@ theorem continuous_symm : Continuous (symm : Path x y → Path y x) := theorem continuous_uncurry_extend_of_continuous_family {ι : Type*} [TopologicalSpace ι] {a b : ι → X} (γ : ∀ t : ι, Path (a t) (b t)) (h : Continuous ↿γ) : Continuous ↿fun t => (γ t).extend := by - apply h.comp (continuous_id.prod_map continuous_projIcc) + apply h.comp (continuous_id.prodMap continuous_projIcc) exact zero_le_one @[continuity] @@ -447,12 +447,12 @@ theorem trans_continuous_family {ι : Type*} [TopologicalSpace ι] refine Continuous.if_le ?_ ?_ (continuous_subtype_val.comp continuous_snd) continuous_const ?_ · change Continuous ((fun p : ι × ℝ => (γ₁ p.1).extend p.2) ∘ Prod.map id (fun x => 2 * x : I → ℝ)) - exact h₁'.comp (continuous_id.prod_map <| continuous_const.mul continuous_subtype_val) + exact h₁'.comp (continuous_id.prodMap <| continuous_const.mul continuous_subtype_val) · change Continuous ((fun p : ι × ℝ => (γ₂ p.1).extend p.2) ∘ Prod.map id (fun x => 2 * x - 1 : I → ℝ)) exact h₂'.comp - (continuous_id.prod_map <| + (continuous_id.prodMap <| (continuous_const.mul continuous_subtype_val).sub continuous_const) · rintro st hst simp [hst, mul_inv_cancel₀ (two_ne_zero' ℝ)] diff --git a/Mathlib/Topology/Constructions.lean b/Mathlib/Topology/Constructions.lean index 803aee5a4888b..655e49f24bd70 100644 --- a/Mathlib/Topology/Constructions.lean +++ b/Mathlib/Topology/Constructions.lean @@ -400,10 +400,12 @@ theorem Continuous.comp₄ {g : X × Y × Z × ζ → ε} (hg : Continuous g) {e hg.comp₃ he hf <| hk.prod_mk hl @[continuity] -theorem Continuous.prod_map {f : Z → X} {g : W → Y} (hf : Continuous f) (hg : Continuous g) : - Continuous fun p : Z × W => (f p.1, g p.2) := +theorem Continuous.prodMap {f : Z → X} {g : W → Y} (hf : Continuous f) (hg : Continuous g) : + Continuous (Prod.map f g) := hf.fst'.prod_mk hg.snd' +@[deprecated (since := "2024-10-05")] alias Continuous.prod_map := Continuous.prodMap + /-- A version of `continuous_inf_dom_left` for binary functions -/ theorem continuous_inf_dom_left₂ {X Y Z} {f : X → Y → Z} {ta1 ta2 : TopologicalSpace X} {tb1 tb2 : TopologicalSpace Y} {tc1 : TopologicalSpace Z} @@ -411,7 +413,7 @@ theorem continuous_inf_dom_left₂ {X Y Z} {f : X → Y → Z} {ta1 ta2 : Topolo haveI := ta1 ⊓ ta2; haveI := tb1 ⊓ tb2; exact Continuous fun p : X × Y => f p.1 p.2 := by have ha := @continuous_inf_dom_left _ _ id ta1 ta2 ta1 (@continuous_id _ (id _)) have hb := @continuous_inf_dom_left _ _ id tb1 tb2 tb1 (@continuous_id _ (id _)) - have h_continuous_id := @Continuous.prod_map _ _ _ _ ta1 tb1 (ta1 ⊓ ta2) (tb1 ⊓ tb2) _ _ ha hb + have h_continuous_id := @Continuous.prodMap _ _ _ _ ta1 tb1 (ta1 ⊓ ta2) (tb1 ⊓ tb2) _ _ ha hb exact @Continuous.comp _ _ _ (id _) (id _) _ _ _ h h_continuous_id /-- A version of `continuous_inf_dom_right` for binary functions -/ @@ -421,7 +423,7 @@ theorem continuous_inf_dom_right₂ {X Y Z} {f : X → Y → Z} {ta1 ta2 : Topol haveI := ta1 ⊓ ta2; haveI := tb1 ⊓ tb2; exact Continuous fun p : X × Y => f p.1 p.2 := by have ha := @continuous_inf_dom_right _ _ id ta1 ta2 ta2 (@continuous_id _ (id _)) have hb := @continuous_inf_dom_right _ _ id tb1 tb2 tb2 (@continuous_id _ (id _)) - have h_continuous_id := @Continuous.prod_map _ _ _ _ ta2 tb2 (ta1 ⊓ ta2) (tb1 ⊓ tb2) _ _ ha hb + have h_continuous_id := @Continuous.prodMap _ _ _ _ ta2 tb2 (ta1 ⊓ ta2) (tb1 ⊓ tb2) _ _ ha hb exact @Continuous.comp _ _ _ (id _) (id _) _ _ _ h h_continuous_id /-- A version of `continuous_sInf_dom` for binary functions -/ @@ -433,7 +435,7 @@ theorem continuous_sInf_dom₂ {X Y Z} {f : X → Y → Z} {tas : Set (Topologic exact @Continuous _ _ _ tc fun p : X × Y => f p.1 p.2 := by have hX := continuous_sInf_dom hX continuous_id have hY := continuous_sInf_dom hY continuous_id - have h_continuous_id := @Continuous.prod_map _ _ _ _ tX tY (sInf tas) (sInf tbs) _ _ hX hY + have h_continuous_id := @Continuous.prodMap _ _ _ _ tX tY (sInf tas) (sInf tbs) _ _ hX hY exact @Continuous.comp _ _ _ (id _) (id _) _ _ _ hf h_continuous_id theorem Filter.Eventually.prod_inl_nhds {p : X → Prop} {x : X} (h : ∀ᶠ x in 𝓝 x, p x) (y : Y) : @@ -582,13 +584,19 @@ theorem ContinuousAt.prod {f : X → Y} {g : X → Z} {x : X} (hf : ContinuousAt (hg : ContinuousAt g x) : ContinuousAt (fun x => (f x, g x)) x := hf.prod_mk_nhds hg -theorem ContinuousAt.prod_map {f : X → Z} {g : Y → W} {p : X × Y} (hf : ContinuousAt f p.fst) - (hg : ContinuousAt g p.snd) : ContinuousAt (fun p : X × Y => (f p.1, g p.2)) p := +theorem ContinuousAt.prodMap {f : X → Z} {g : Y → W} {p : X × Y} (hf : ContinuousAt f p.fst) + (hg : ContinuousAt g p.snd) : ContinuousAt (Prod.map f g) p := hf.fst''.prod hg.snd'' -theorem ContinuousAt.prod_map' {f : X → Z} {g : Y → W} {x : X} {y : Y} (hf : ContinuousAt f x) - (hg : ContinuousAt g y) : ContinuousAt (fun p : X × Y => (f p.1, g p.2)) (x, y) := - hf.fst'.prod hg.snd' +@[deprecated (since := "2024-10-05")] alias ContinuousAt.prod_map := ContinuousAt.prodMap + +/-- A version of `ContinuousAt.prodMap` that avoids `Prod.fst`/`Prod.snd` +by assuming that the point is `(x, y)`. -/ +theorem ContinuousAt.prodMap' {f : X → Z} {g : Y → W} {x : X} {y : Y} (hf : ContinuousAt f x) + (hg : ContinuousAt g y) : ContinuousAt (Prod.map f g) (x, y) := + hf.prodMap hg + +@[deprecated (since := "2024-10-05")] alias ContinuousAt.prod_map' := ContinuousAt.prodMap' theorem ContinuousAt.comp₂ {f : Y × Z → W} {g : X → Y} {h : X → Z} {x : X} (hf : ContinuousAt f (g x, h x)) (hg : ContinuousAt g x) (hh : ContinuousAt h x) : @@ -776,15 +784,19 @@ theorem Dense.prod {s : Set X} {t : Set Y} (hs : Dense s) (ht : Dense t) : Dense exact ⟨hs x.1, ht x.2⟩ /-- If `f` and `g` are maps with dense range, then `Prod.map f g` has dense range. -/ -theorem DenseRange.prod_map {ι : Type*} {κ : Type*} {f : ι → Y} {g : κ → Z} (hf : DenseRange f) +theorem DenseRange.prodMap {ι : Type*} {κ : Type*} {f : ι → Y} {g : κ → Z} (hf : DenseRange f) (hg : DenseRange g) : DenseRange (Prod.map f g) := by simpa only [DenseRange, prod_range_range_eq] using hf.prod hg -theorem Inducing.prod_map {f : X → Y} {g : Z → W} (hf : Inducing f) (hg : Inducing g) : +@[deprecated (since := "2024-10-05")] alias DenseRange.prod_map := DenseRange.prodMap + +theorem Inducing.prodMap {f : X → Y} {g : Z → W} (hf : Inducing f) (hg : Inducing g) : Inducing (Prod.map f g) := inducing_iff_nhds.2 fun (x, z) => by simp_rw [Prod.map_def, nhds_prod_eq, hf.nhds_eq_comap, hg.nhds_eq_comap, prod_comap_comap_eq] +@[deprecated (since := "2024-10-05")] alias Inducing.prod_map := Inducing.prodMap + @[simp] theorem inducing_const_prod {x : X} {f : Y → Z} : (Inducing fun x' => (x, f x')) ↔ Inducing f := by simp_rw [inducing_iff, instTopologicalSpaceProd, induced_inf, induced_compose, Function.comp_def, @@ -795,21 +807,27 @@ theorem inducing_prod_const {y : Y} {f : X → Z} : (Inducing fun x => (f x, y)) simp_rw [inducing_iff, instTopologicalSpaceProd, induced_inf, induced_compose, Function.comp_def, induced_const, inf_top_eq] -theorem Embedding.prod_map {f : X → Y} {g : Z → W} (hf : Embedding f) (hg : Embedding g) : +theorem Embedding.prodMap {f : X → Y} {g : Z → W} (hf : Embedding f) (hg : Embedding g) : Embedding (Prod.map f g) := - { hf.toInducing.prod_map hg.toInducing with + { hf.toInducing.prodMap hg.toInducing with inj := fun ⟨x₁, z₁⟩ ⟨x₂, z₂⟩ => by simp [hf.inj.eq_iff, hg.inj.eq_iff] } -protected theorem IsOpenMap.prod {f : X → Y} {g : Z → W} (hf : IsOpenMap f) (hg : IsOpenMap g) : - IsOpenMap fun p : X × Z => (f p.1, g p.2) := by +@[deprecated (since := "2024-10-05")] alias Embedding.prod_map := Embedding.prodMap + +protected theorem IsOpenMap.prodMap {f : X → Y} {g : Z → W} (hf : IsOpenMap f) (hg : IsOpenMap g) : + IsOpenMap (Prod.map f g) := by rw [isOpenMap_iff_nhds_le] rintro ⟨a, b⟩ - rw [nhds_prod_eq, nhds_prod_eq, ← Filter.prod_map_map_eq] + rw [nhds_prod_eq, nhds_prod_eq, ← Filter.prod_map_map_eq'] exact Filter.prod_mono (hf.nhds_le a) (hg.nhds_le b) -protected theorem OpenEmbedding.prod {f : X → Y} {g : Z → W} (hf : OpenEmbedding f) - (hg : OpenEmbedding g) : OpenEmbedding fun x : X × Z => (f x.1, g x.2) := - openEmbedding_of_embedding_open (hf.1.prod_map hg.1) (hf.isOpenMap.prod hg.isOpenMap) +@[deprecated (since := "2024-10-05")] alias IsOpenMap.prod := IsOpenMap.prodMap + +protected theorem OpenEmbedding.prodMap {f : X → Y} {g : Z → W} (hf : OpenEmbedding f) + (hg : OpenEmbedding g) : OpenEmbedding (Prod.map f g) := + openEmbedding_of_embedding_open (hf.1.prodMap hg.1) (hf.isOpenMap.prodMap hg.isOpenMap) + +@[deprecated (since := "2024-10-05")] alias OpenEmbedding.prod := OpenEmbedding.prodMap theorem embedding_graph {f : X → Y} (hf : Continuous f) : Embedding fun x => (x, f x) := embedding_of_embedding_compose (continuous_id.prod_mk hf) continuous_fst embedding_id @@ -819,7 +837,7 @@ theorem embedding_prod_mk (x : X) : Embedding (Prod.mk x : Y → X × Y) := theorem IsOpenQuotientMap.prodMap {f : X → Y} {g : Z → W} (hf : IsOpenQuotientMap f) (hg : IsOpenQuotientMap g) : IsOpenQuotientMap (Prod.map f g) := - ⟨.prodMap hf.1 hg.1, .prod_map hf.2 hg.2, .prod hf.3 hg.3⟩ + ⟨.prodMap hf.1 hg.1, .prodMap hf.2 hg.2, .prodMap hf.3 hg.3⟩ end Prod diff --git a/Mathlib/Topology/ContinuousMap/Algebra.lean b/Mathlib/Topology/ContinuousMap/Algebra.lean index a31ec711b86dc..4e62f17e049be 100644 --- a/Mathlib/Topology/ContinuousMap/Algebra.lean +++ b/Mathlib/Topology/ContinuousMap/Algebra.lean @@ -295,9 +295,9 @@ instance [LocallyCompactSpace α] [Mul β] [ContinuousMul β] : ContinuousMul C( ⟨by refine continuous_of_continuous_uncurry _ ?_ have h1 : Continuous fun x : (C(α, β) × C(α, β)) × α => x.fst.fst x.snd := - continuous_eval.comp (continuous_fst.prod_map continuous_id) + continuous_eval.comp (continuous_fst.prodMap continuous_id) have h2 : Continuous fun x : (C(α, β) × C(α, β)) × α => x.fst.snd x.snd := - continuous_eval.comp (continuous_snd.prod_map continuous_id) + continuous_eval.comp (continuous_snd.prodMap continuous_id) exact h1.mul h2⟩ /-- Coercion to a function as a `MonoidHom`. Similar to `MonoidHom.coeFn`. -/ @@ -550,7 +550,7 @@ instance [LocallyCompactSpace α] [TopologicalSpace R] [SMul R M] [ContinuousSMu ⟨by refine continuous_of_continuous_uncurry _ ?_ have h : Continuous fun x : (R × C(α, M)) × α => x.fst.snd x.snd := - continuous_eval.comp (continuous_snd.prod_map continuous_id) + continuous_eval.comp (continuous_snd.prodMap continuous_id) exact (continuous_fst.comp continuous_fst).smul h⟩ @[to_additive (attr := simp, norm_cast)] diff --git a/Mathlib/Topology/ContinuousMap/Basic.lean b/Mathlib/Topology/ContinuousMap/Basic.lean index 34583c8f81047..c5f501967a7b2 100644 --- a/Mathlib/Topology/ContinuousMap/Basic.lean +++ b/Mathlib/Topology/ContinuousMap/Basic.lean @@ -271,8 +271,6 @@ def prodMk (f : C(α, β₁)) (g : C(α, β₂)) : C(α, β₁ × β₂) where @[simps] def prodMap (f : C(α₁, α₂)) (g : C(β₁, β₂)) : C(α₁ × β₁, α₂ × β₂) where toFun := Prod.map f g - continuous_toFun := f.continuous.prod_map g.continuous - -- Porting note: proof was `continuity` @[simp] theorem prod_eval (f : C(α, β₁)) (g : C(α, β₂)) (a : α) : (prodMk f g) a = (f a, g a) := diff --git a/Mathlib/Topology/DenseEmbedding.lean b/Mathlib/Topology/DenseEmbedding.lean index 84bfa9ec96c78..210caa3dca529 100644 --- a/Mathlib/Topology/DenseEmbedding.lean +++ b/Mathlib/Topology/DenseEmbedding.lean @@ -81,11 +81,14 @@ theorem interior_compact_eq_empty [T2Space β] (di : IsDenseInducing i) (hd : De exact hyi (image_subset_range _ _ hys) /-- The product of two dense inducings is a dense inducing -/ -protected theorem prod [TopologicalSpace γ] [TopologicalSpace δ] {e₁ : α → β} {e₂ : γ → δ} +protected theorem prodMap [TopologicalSpace γ] [TopologicalSpace δ] {e₁ : α → β} {e₂ : γ → δ} (de₁ : IsDenseInducing e₁) (de₂ : IsDenseInducing e₂) : - IsDenseInducing fun p : α × γ => (e₁ p.1, e₂ p.2) where - toInducing := de₁.toInducing.prod_map de₂.toInducing - dense := de₁.dense.prod_map de₂.dense + IsDenseInducing (Prod.map e₁ e₂) where + toInducing := de₁.toInducing.prodMap de₂.toInducing + dense := de₁.dense.prodMap de₂.dense + +@[deprecated (since := "2024-10-06")] +protected alias prod := IsDenseInducing.prodMap open TopologicalSpace @@ -234,11 +237,13 @@ protected theorem separableSpace [SeparableSpace α] (de : IsDenseEmbedding e) : de.toIsDenseInducing.separableSpace /-- The product of two dense embeddings is a dense embedding. -/ -protected theorem prod {e₁ : α → β} {e₂ : γ → δ} (de₁ : IsDenseEmbedding e₁) +protected theorem prodMap {e₁ : α → β} {e₂ : γ → δ} (de₁ : IsDenseEmbedding e₁) (de₂ : IsDenseEmbedding e₂) : IsDenseEmbedding fun p : α × γ => (e₁ p.1, e₂ p.2) := - { de₁.toIsDenseInducing.prod de₂.toIsDenseInducing with + { de₁.toIsDenseInducing.prodMap de₂.toIsDenseInducing with inj := de₁.inj.prodMap de₂.inj } +@[deprecated (since := "2024-10-06")] protected alias prod := IsDenseEmbedding.prodMap + /-- The dense embedding of a subtype inside its closure. -/ @[simps] def subtypeEmb {α : Type*} (p : α → Prop) (e : α → β) (x : { x // p x }) : @@ -288,14 +293,14 @@ theorem isClosed_property [TopologicalSpace β] {e : α → β} {p : β → Prop theorem isClosed_property2 [TopologicalSpace β] {e : α → β} {p : β → β → Prop} (he : DenseRange e) (hp : IsClosed { q : β × β | p q.1 q.2 }) (h : ∀ a₁ a₂, p (e a₁) (e a₂)) : ∀ b₁ b₂, p b₁ b₂ := - have : ∀ q : β × β, p q.1 q.2 := isClosed_property (he.prod_map he) hp fun _ => h _ _ + have : ∀ q : β × β, p q.1 q.2 := isClosed_property (he.prodMap he) hp fun _ => h _ _ fun b₁ b₂ => this ⟨b₁, b₂⟩ theorem isClosed_property3 [TopologicalSpace β] {e : α → β} {p : β → β → β → Prop} (he : DenseRange e) (hp : IsClosed { q : β × β × β | p q.1 q.2.1 q.2.2 }) (h : ∀ a₁ a₂ a₃, p (e a₁) (e a₂) (e a₃)) : ∀ b₁ b₂ b₃, p b₁ b₂ b₃ := have : ∀ q : β × β × β, p q.1 q.2.1 q.2.2 := - isClosed_property (he.prod_map <| he.prod_map he) hp fun _ => h _ _ _ + isClosed_property (he.prodMap <| he.prodMap he) hp fun _ => h _ _ _ fun b₁ b₂ b₃ => this ⟨b₁, b₂, b₃⟩ @[elab_as_elim] diff --git a/Mathlib/Topology/FiberBundle/Constructions.lean b/Mathlib/Topology/FiberBundle/Constructions.lean index 8a7ea9e221f3c..b5cac2629b8cf 100644 --- a/Mathlib/Topology/FiberBundle/Constructions.lean +++ b/Mathlib/Topology/FiberBundle/Constructions.lean @@ -138,8 +138,7 @@ theorem Prod.continuous_to_fun : ContinuousOn (Prod.toFun' e₁ e₂) have hf₁ : Continuous f₁ := (Prod.inducing_diag F₁ E₁ F₂ E₂).continuous have hf₂ : ContinuousOn f₂ (e₁.source ×ˢ e₂.source) := e₁.toPartialHomeomorph.continuousOn.prod_map e₂.toPartialHomeomorph.continuousOn - have hf₃ : Continuous f₃ := - (continuous_fst.comp continuous_fst).prod_mk (continuous_snd.prod_map continuous_snd) + have hf₃ : Continuous f₃ := by fun_prop refine ((hf₃.comp_continuousOn hf₂).comp hf₁.continuousOn ?_).congr ?_ · rw [e₁.source_eq, e₂.source_eq] exact mapsTo_preimage _ _ @@ -176,8 +175,7 @@ theorem Prod.right_inv {x : B × F₁ × F₂} theorem Prod.continuous_inv_fun : ContinuousOn (Prod.invFun' e₁ e₂) ((e₁.baseSet ∩ e₂.baseSet) ×ˢ univ) := by rw [(Prod.inducing_diag F₁ E₁ F₂ E₂).continuousOn_iff] - have H₁ : Continuous fun p : B × F₁ × F₂ ↦ ((p.1, p.2.1), (p.1, p.2.2)) := - (continuous_id.prod_map continuous_fst).prod_mk (continuous_id.prod_map continuous_snd) + have H₁ : Continuous fun p : B × F₁ × F₂ ↦ ((p.1, p.2.1), (p.1, p.2.2)) := by fun_prop refine (e₁.continuousOn_symm.prod_map e₂.continuousOn_symm).comp H₁.continuousOn ?_ exact fun x h ↦ ⟨⟨h.1.1, mem_univ _⟩, ⟨h.1.2, mem_univ _⟩⟩ @@ -226,7 +224,7 @@ variable [∀ x, Zero (E₁ x)] [∀ x, Zero (E₂ x)] [∀ x : B, TopologicalSp noncomputable instance FiberBundle.prod : FiberBundle (F₁ × F₂) (E₁ ×ᵇ E₂) where totalSpaceMk_inducing' b := by rw [← (Prod.inducing_diag F₁ E₁ F₂ E₂).of_comp_iff] - exact (totalSpaceMk_inducing F₁ E₁ b).prod_map (totalSpaceMk_inducing F₂ E₂ b) + exact (totalSpaceMk_inducing F₁ E₁ b).prodMap (totalSpaceMk_inducing F₂ E₂ b) trivializationAtlas' := { e | ∃ (e₁ : Trivialization F₁ (π F₁ E₁)) (e₂ : Trivialization F₂ (π F₂ E₂)) (_ : MemTrivializationAtlas e₁) (_ : MemTrivializationAtlas e₂), @@ -334,7 +332,7 @@ noncomputable def Trivialization.pullback (e : Trivialization F (π F E)) (f : K pullbackTotalSpaceEmbedding] refine continuousOn_fst.prod - (e.continuousOn_symm.comp ((map_continuous f).prod_map continuous_id).continuousOn + (e.continuousOn_symm.comp ((map_continuous f).prodMap continuous_id).continuousOn Subset.rfl) source_eq := by dsimp only diff --git a/Mathlib/Topology/Homeomorph.lean b/Mathlib/Topology/Homeomorph.lean index bcb1cd5d76a91..6d33a57317001 100644 --- a/Mathlib/Topology/Homeomorph.lean +++ b/Mathlib/Topology/Homeomorph.lean @@ -498,8 +498,6 @@ def sumCongr (h₁ : X ≃ₜ X') (h₂ : Y ≃ₜ Y') : X ⊕ Y ≃ₜ X' ⊕ Y /-- Product of two homeomorphisms. -/ def prodCongr (h₁ : X ≃ₜ X') (h₂ : Y ≃ₜ Y') : X × Y ≃ₜ X' × Y' where - continuous_toFun := h₁.continuous.prod_map h₂.continuous - continuous_invFun := h₁.symm.continuous.prod_map h₂.symm.continuous toEquiv := h₁.toEquiv.prodCongr h₂.toEquiv @[simp] @@ -696,9 +694,9 @@ section Distrib def sumProdDistrib : (X ⊕ Y) × Z ≃ₜ (X × Z) ⊕ (Y × Z) := Homeomorph.symm <| homeomorphOfContinuousOpen (Equiv.sumProdDistrib X Y Z).symm - ((continuous_inl.prod_map continuous_id).sum_elim - (continuous_inr.prod_map continuous_id)) <| - (isOpenMap_inl.prod IsOpenMap.id).sum_elim (isOpenMap_inr.prod IsOpenMap.id) + ((continuous_inl.prodMap continuous_id).sum_elim + (continuous_inr.prodMap continuous_id)) <| + (isOpenMap_inl.prodMap IsOpenMap.id).sum_elim (isOpenMap_inr.prodMap IsOpenMap.id) /-- `X × (Y ⊕ Z)` is homeomorphic to `X × Y ⊕ X × Z`. -/ def prodSumDistrib : X × (Y ⊕ Z) ≃ₜ (X × Y) ⊕ (X × Z) := @@ -712,7 +710,7 @@ def sigmaProdDistrib : (Σ i, X i) × Y ≃ₜ Σ i, X i × Y := Homeomorph.symm <| homeomorphOfContinuousOpen (Equiv.sigmaProdDistrib X Y).symm (continuous_sigma fun _ => continuous_sigmaMk.fst'.prod_mk continuous_snd) - (isOpenMap_sigma.2 fun _ => isOpenMap_sigmaMk.prod IsOpenMap.id) + (isOpenMap_sigma.2 fun _ => isOpenMap_sigmaMk.prodMap IsOpenMap.id) end Distrib @@ -952,7 +950,7 @@ lemma IsHomeomorph.sumMap {g : Z → W} (hf : IsHomeomorph f) (hg : IsHomeomorph IsHomeomorph (Sum.map f g) := ⟨hf.1.sum_map hg.1, hf.2.sumMap hg.2, hf.3.sum_map hg.3⟩ lemma IsHomeomorph.prodMap {g : Z → W} (hf : IsHomeomorph f) (hg : IsHomeomorph g) : - IsHomeomorph (Prod.map f g) := ⟨hf.1.prod_map hg.1, hf.2.prod hg.2, hf.3.prodMap hg.3⟩ + IsHomeomorph (Prod.map f g) := ⟨hf.1.prodMap hg.1, hf.2.prodMap hg.2, hf.3.prodMap hg.3⟩ lemma IsHomeomorph.sigmaMap {ι κ : Type*} {X : ι → Type*} {Y : κ → Type*} [∀ i, TopologicalSpace (X i)] [∀ i, TopologicalSpace (Y i)] {f : ι → κ} diff --git a/Mathlib/Topology/Homotopy/HomotopyGroup.lean b/Mathlib/Topology/Homotopy/HomotopyGroup.lean index 93fad3a1199db..daaa41db02500 100644 --- a/Mathlib/Topology/Homotopy/HomotopyGroup.lean +++ b/Mathlib/Topology/Homotopy/HomotopyGroup.lean @@ -185,7 +185,7 @@ theorem continuous_toLoop (i : N) : Continuous (@toLoop N X _ x _ i) := Path.continuous_uncurry_iff.1 <| Continuous.subtype_mk (ContinuousMap.continuous_eval.comp <| - Continuous.prod_map + Continuous.prodMap (ContinuousMap.continuous_curry.comp <| (ContinuousMap.continuous_comp_left _).comp continuous_subtype_val) continuous_id) diff --git a/Mathlib/Topology/Inseparable.lean b/Mathlib/Topology/Inseparable.lean index f122cb8067af0..a179072f0d072 100644 --- a/Mathlib/Topology/Inseparable.lean +++ b/Mathlib/Topology/Inseparable.lean @@ -5,6 +5,7 @@ Authors: Andrew Yang, Yury Kudryashov -/ import Mathlib.Tactic.TFAE import Mathlib.Topology.ContinuousOn +import Mathlib.Topology.Maps.OpenQuotient /-! # Inseparable points in a topological space @@ -602,14 +603,8 @@ theorem map_mk_nhdsWithin_preimage (s : Set (SeparationQuotient X)) (x : X) : rw [nhdsWithin, ← comap_principal, Filter.push_pull, nhdsWithin, map_mk_nhds] /-- The map `(x, y) ↦ (mk x, mk y)` is a quotient map. -/ -theorem quotientMap_prodMap_mk : QuotientMap (Prod.map mk mk : X × Y → _) := by - have hsurj : Surjective (Prod.map mk mk : X × Y → _) := surjective_mk.prodMap surjective_mk - refine quotientMap_iff.2 ⟨hsurj, fun s ↦ ?_⟩ - refine ⟨fun hs ↦ hs.preimage (continuous_mk.prod_map continuous_mk), fun hs ↦ ?_⟩ - refine isOpen_iff_mem_nhds.2 <| hsurj.forall.2 fun (x, y) h ↦ ?_ - rw [Prod.map_mk, nhds_prod_eq, ← map_mk_nhds, ← map_mk_nhds, Filter.prod_map_map_eq', - ← nhds_prod_eq, Filter.mem_map] - exact hs.mem_nhds h +theorem quotientMap_prodMap_mk : QuotientMap (Prod.map mk mk : X × Y → _) := + (isOpenQuotientMap_mk.prodMap isOpenQuotientMap_mk).quotientMap /-- Lift a map `f : X → α` such that `Inseparable x y → f x = f y` to a map `SeparationQuotient X → α`. -/ diff --git a/Mathlib/Topology/Instances/ENNReal.lean b/Mathlib/Topology/Instances/ENNReal.lean index 609c01f853877..5623966b95601 100644 --- a/Mathlib/Topology/Instances/ENNReal.lean +++ b/Mathlib/Topology/Instances/ENNReal.lean @@ -91,7 +91,7 @@ theorem continuousAt_coe_iff {α : Type*} [TopologicalSpace α] {x : ℝ≥0} {f theorem nhds_coe_coe {r p : ℝ≥0} : 𝓝 ((r : ℝ≥0∞), (p : ℝ≥0∞)) = (𝓝 (r, p)).map fun p : ℝ≥0 × ℝ≥0 => (↑p.1, ↑p.2) := - ((openEmbedding_coe.prod openEmbedding_coe).map_nhds_eq (r, p)).symm + ((openEmbedding_coe.prodMap openEmbedding_coe).map_nhds_eq (r, p)).symm theorem continuous_ofReal : Continuous ENNReal.ofReal := (continuous_coe_iff.2 continuous_id).comp continuous_real_toNNReal diff --git a/Mathlib/Topology/Instances/EReal.lean b/Mathlib/Topology/Instances/EReal.lean index 8b4ae160970e4..4d970aa153cb8 100644 --- a/Mathlib/Topology/Instances/EReal.lean +++ b/Mathlib/Topology/Instances/EReal.lean @@ -70,7 +70,7 @@ theorem nhds_coe {r : ℝ} : 𝓝 (r : EReal) = (𝓝 r).map (↑) := theorem nhds_coe_coe {r p : ℝ} : 𝓝 ((r : EReal), (p : EReal)) = (𝓝 (r, p)).map fun p : ℝ × ℝ => (↑p.1, ↑p.2) := - ((openEmbedding_coe.prod openEmbedding_coe).map_nhds_eq (r, p)).symm + ((openEmbedding_coe.prodMap openEmbedding_coe).map_nhds_eq (r, p)).symm theorem tendsto_toReal {a : EReal} (ha : a ≠ ⊤) (h'a : a ≠ ⊥) : Tendsto EReal.toReal (𝓝 a) (𝓝 a.toReal) := by @@ -358,7 +358,7 @@ private lemma continuousAt_mul_symm1 {a b : EReal} simp rw [this] apply ContinuousAt.comp (Continuous.continuousAt continuous_neg) - <| ContinuousAt.comp _ (ContinuousAt.prod_map (Continuous.continuousAt continuous_neg) + <| ContinuousAt.comp _ (ContinuousAt.prodMap (Continuous.continuousAt continuous_neg) (Continuous.continuousAt continuous_id)) simp [h] diff --git a/Mathlib/Topology/Instances/Rat.lean b/Mathlib/Topology/Instances/Rat.lean index a5585543b41fe..c6e661d02abb8 100644 --- a/Mathlib/Topology/Instances/Rat.lean +++ b/Mathlib/Topology/Instances/Rat.lean @@ -84,7 +84,7 @@ theorem uniformContinuous_add : UniformContinuous fun p : ℚ × ℚ => p.1 + p. Rat.isUniformEmbedding_coe_real.toUniformInducing.uniformContinuous_iff.2 <| by simp only [Function.comp_def, Rat.cast_add] exact Real.uniformContinuous_add.comp - (Rat.uniformContinuous_coe_real.prod_map Rat.uniformContinuous_coe_real) + (Rat.uniformContinuous_coe_real.prodMap Rat.uniformContinuous_coe_real) theorem uniformContinuous_neg : UniformContinuous (@Neg.neg ℚ _) := Metric.uniformContinuous_iff.2 fun ε ε0 => diff --git a/Mathlib/Topology/Maps/Proper/Basic.lean b/Mathlib/Topology/Maps/Proper/Basic.lean index 4132dd33262f6..00bd0a3bd8677 100644 --- a/Mathlib/Topology/Maps/Proper/Basic.lean +++ b/Mathlib/Topology/Maps/Proper/Basic.lean @@ -174,12 +174,12 @@ lemma isProperMap_of_comp_of_t2 [T2Space Y] (hf : Continuous f) (hg : Continuous exact ⟨x, hx⟩ /-- A binary product of proper maps is proper. -/ -lemma IsProperMap.prod_map {g : Z → W} (hf : IsProperMap f) (hg : IsProperMap g) : +lemma IsProperMap.prodMap {g : Z → W} (hf : IsProperMap f) (hg : IsProperMap g) : IsProperMap (Prod.map f g) := by simp_rw [isProperMap_iff_ultrafilter] at hf hg ⊢ constructor -- Continuity is clear. - · exact hf.1.prod_map hg.1 + · exact hf.1.prodMap hg.1 -- Let `𝒰 : Ultrafilter (X × Z)`, and assume that `f × g` tends to some `(y, w) : Y × W` -- along `𝒰`. · intro 𝒰 ⟨y, w⟩ hyw @@ -197,6 +197,8 @@ lemma IsProperMap.prod_map {g : Z → W} (hf : IsProperMap f) (hg : IsProperMap rw [nhds_prod_eq, le_prod] exact ⟨hx, hz⟩ +@[deprecated (since := "2024-10-06")] alias IsProperMap.prod_map := IsProperMap.prodMap + /-- Any product of proper maps is proper. -/ lemma IsProperMap.pi_map {X Y : ι → Type*} [∀ i, TopologicalSpace (X i)] [∀ i, TopologicalSpace (Y i)] {f : (i : ι) → X i → Y i} (h : ∀ i, IsProperMap (f i)) : @@ -416,7 +418,7 @@ easier to use because it allows `Z` to live in any universe. -/ theorem IsProperMap.universally_closed (Z) [TopologicalSpace Z] (h : IsProperMap f) : IsClosedMap (Prod.map f id : X × Z → Y × Z) := -- `f × id` is proper as a product of proper maps, hence closed. - (h.prod_map isProperMap_id).isClosedMap + (h.prodMap isProperMap_id).isClosedMap /-- A map `f : X → Y` is proper if and only if it is continuous and the map `(Prod.map f id : X × Filter X → Y × Filter X)` is closed. This is stronger than diff --git a/Mathlib/Topology/Order/OrderClosed.lean b/Mathlib/Topology/Order/OrderClosed.lean index 7a476daeaf513..e2763d592bc24 100644 --- a/Mathlib/Topology/Order/OrderClosed.lean +++ b/Mathlib/Topology/Order/OrderClosed.lean @@ -543,7 +543,7 @@ namespace Subtype -- todo: add `OrderEmbedding.orderClosedTopology` instance {p : α → Prop} : OrderClosedTopology (Subtype p) := have this : Continuous fun p : Subtype p × Subtype p => ((p.fst : α), (p.snd : α)) := - continuous_subtype_val.prod_map continuous_subtype_val + continuous_subtype_val.prodMap continuous_subtype_val OrderClosedTopology.mk (t.isClosed_le'.preimage this) end Subtype diff --git a/Mathlib/Topology/Separation.lean b/Mathlib/Topology/Separation.lean index 2f7f8cfdcd660..c2de323af42d9 100644 --- a/Mathlib/Topology/Separation.lean +++ b/Mathlib/Topology/Separation.lean @@ -1451,7 +1451,7 @@ theorem Set.InjOn.exists_mem_nhdsSet {X Y : Type*} [TopologicalSpace X] [Topolog · rcases loc x hx with ⟨u, hu, hf⟩ exact Filter.mem_of_superset (prod_mem_nhds hu hu) <| forall_prod_set.2 hf · suffices ∀ᶠ z in 𝓝 (x, y), f z.1 ≠ f z.2 from this.mono fun _ hne h ↦ absurd h hne - refine (fc x hx).prod_map' (fc y hy) <| isClosed_diagonal.isOpen_compl.mem_nhds ?_ + refine (fc x hx).prodMap' (fc y hy) <| isClosed_diagonal.isOpen_compl.mem_nhds ?_ exact inj.ne hx hy hne rw [← eventually_nhdsSet_iff_forall, sc.nhdsSet_prod_eq sc] at this exact eventually_prod_self_iff.1 this diff --git a/Mathlib/Topology/UniformSpace/AbstractCompletion.lean b/Mathlib/Topology/UniformSpace/AbstractCompletion.lean index 9b5cc359931bb..d71b406019079 100644 --- a/Mathlib/Topology/UniformSpace/AbstractCompletion.lean +++ b/Mathlib/Topology/UniformSpace/AbstractCompletion.lean @@ -306,7 +306,7 @@ protected def prod : AbstractCompletion (α × β) where complete := inferInstance separation := inferInstance uniformInducing := UniformInducing.prod pkg.uniformInducing pkg'.uniformInducing - dense := DenseRange.prod_map pkg.dense pkg'.dense + dense := pkg.dense.prodMap pkg'.dense end Prod diff --git a/Mathlib/Topology/UniformSpace/Basic.lean b/Mathlib/Topology/UniformSpace/Basic.lean index d2dc3a7e09b3f..caab9ab1db8cb 100644 --- a/Mathlib/Topology/UniformSpace/Basic.lean +++ b/Mathlib/Topology/UniformSpace/Basic.lean @@ -1464,10 +1464,12 @@ theorem UniformContinuous.prod_mk_right {f : α × β → γ} (h : UniformContin UniformContinuous fun b => f (a, b) := h.comp (uniformContinuous_const.prod_mk uniformContinuous_id) -theorem UniformContinuous.prod_map [UniformSpace δ] {f : α → γ} {g : β → δ} +theorem UniformContinuous.prodMap [UniformSpace δ] {f : α → γ} {g : β → δ} (hf : UniformContinuous f) (hg : UniformContinuous g) : UniformContinuous (Prod.map f g) := (hf.comp uniformContinuous_fst).prod_mk (hg.comp uniformContinuous_snd) +@[deprecated (since := "2024-10-06")] alias UniformContinuous.prod_map := UniformContinuous.prodMap + theorem toTopologicalSpace_prod {α} {β} [u : UniformSpace α] [v : UniformSpace β] : @UniformSpace.toTopologicalSpace (α × β) instUniformSpaceProd = @instTopologicalSpaceProd α β u.toTopologicalSpace v.toTopologicalSpace := @@ -1483,7 +1485,7 @@ theorem uniformContinuous_inf_dom_left₂ {α β γ} {f : α → β → γ} {ua1 have ha := @UniformContinuous.inf_dom_left _ _ id ua1 ua2 ua1 (@uniformContinuous_id _ (id _)) have hb := @UniformContinuous.inf_dom_left _ _ id ub1 ub2 ub1 (@uniformContinuous_id _ (id _)) have h_unif_cont_id := - @UniformContinuous.prod_map _ _ _ _ (ua1 ⊓ ua2) (ub1 ⊓ ub2) ua1 ub1 _ _ ha hb + @UniformContinuous.prodMap _ _ _ _ (ua1 ⊓ ua2) (ub1 ⊓ ub2) ua1 ub1 _ _ ha hb exact @UniformContinuous.comp _ _ _ (id _) (id _) _ _ _ h h_unif_cont_id /-- A version of `UniformContinuous.inf_dom_right` for binary functions -/ @@ -1496,7 +1498,7 @@ theorem uniformContinuous_inf_dom_right₂ {α β γ} {f : α → β → γ} {ua have ha := @UniformContinuous.inf_dom_right _ _ id ua1 ua2 ua2 (@uniformContinuous_id _ (id _)) have hb := @UniformContinuous.inf_dom_right _ _ id ub1 ub2 ub2 (@uniformContinuous_id _ (id _)) have h_unif_cont_id := - @UniformContinuous.prod_map _ _ _ _ (ua1 ⊓ ua2) (ub1 ⊓ ub2) ua2 ub2 _ _ ha hb + @UniformContinuous.prodMap _ _ _ _ (ua1 ⊓ ua2) (ub1 ⊓ ub2) ua2 ub2 _ _ ha hb exact @UniformContinuous.comp _ _ _ (id _) (id _) _ _ _ h h_unif_cont_id /-- A version of `uniformContinuous_sInf_dom` for binary functions -/ @@ -1509,7 +1511,7 @@ theorem uniformContinuous_sInf_dom₂ {α β γ} {f : α → β → γ} {uas : S let _ : UniformSpace (α × β) := instUniformSpaceProd have ha := uniformContinuous_sInf_dom ha uniformContinuous_id have hb := uniformContinuous_sInf_dom hb uniformContinuous_id - have h_unif_cont_id := @UniformContinuous.prod_map _ _ _ _ (sInf uas) (sInf ubs) ua ub _ _ ha hb + have h_unif_cont_id := @UniformContinuous.prodMap _ _ _ _ (sInf uas) (sInf ubs) ua ub _ _ ha hb exact @UniformContinuous.comp _ _ _ (id _) (id _) _ _ _ hf h_unif_cont_id end Prod @@ -1545,7 +1547,7 @@ theorem UniformContinuous₂.comp {f : α → β → γ} {g : γ → δ} (hg : U theorem UniformContinuous₂.bicompl {f : α → β → γ} {ga : δ → α} {gb : δ' → β} (hf : UniformContinuous₂ f) (hga : UniformContinuous ga) (hgb : UniformContinuous gb) : UniformContinuous₂ (bicompl f ga gb) := - hf.uniformContinuous.comp (hga.prod_map hgb) + hf.uniformContinuous.comp (hga.prodMap hgb) end @@ -1737,7 +1739,7 @@ theorem continuousAt_iff'_left [TopologicalSpace β] {f : β → α} {b : β} : theorem continuousAt_iff_prod [TopologicalSpace β] {f : β → α} {b : β} : ContinuousAt f b ↔ Tendsto (fun x : β × β => (f x.1, f x.2)) (𝓝 (b, b)) (𝓤 α) := - ⟨fun H => le_trans (H.prod_map' H) (nhds_le_uniformity _), fun H => + ⟨fun H => le_trans (H.prodMap' H) (nhds_le_uniformity _), fun H => continuousAt_iff'_left.2 <| H.comp <| tendsto_id.prod_mk_nhds tendsto_const_nhds⟩ theorem continuousWithinAt_iff'_right [TopologicalSpace β] {f : β → α} {b : β} {s : Set β} : diff --git a/Mathlib/Topology/UniformSpace/Compact.lean b/Mathlib/Topology/UniformSpace/Compact.lean index 8f4de0a668e02..175d8d75ed946 100644 --- a/Mathlib/Topology/UniformSpace/Compact.lean +++ b/Mathlib/Topology/UniformSpace/Compact.lean @@ -151,7 +151,7 @@ theorem CompactSpace.uniformContinuous_of_continuous [CompactSpace α] {f : α (h : Continuous f) : UniformContinuous f := calc map (Prod.map f f) (𝓤 α) = map (Prod.map f f) (𝓝ˢ (diagonal α)) := by rw [nhdsSet_diagonal_eq_uniformity] - _ ≤ 𝓝ˢ (diagonal β) := (h.prod_map h).tendsto_nhdsSet mapsTo_prod_map_diagonal + _ ≤ 𝓝ˢ (diagonal β) := (h.prodMap h).tendsto_nhdsSet mapsTo_prod_map_diagonal _ ≤ 𝓤 β := nhdsSet_diagonal_le_uniformity /-- Heine-Cantor: a continuous function on a compact set of a uniform space is uniformly diff --git a/Mathlib/Topology/UniformSpace/Completion.lean b/Mathlib/Topology/UniformSpace/Completion.lean index 2ca3f3f72a834..b8973237d30ae 100644 --- a/Mathlib/Topology/UniformSpace/Completion.lean +++ b/Mathlib/Topology/UniformSpace/Completion.lean @@ -394,12 +394,12 @@ alias denseEmbedding_coe := isDenseEmbedding_coe theorem denseRange_coe₂ : DenseRange fun x : α × β => ((x.1 : Completion α), (x.2 : Completion β)) := - denseRange_coe.prod_map denseRange_coe + denseRange_coe.prodMap denseRange_coe theorem denseRange_coe₃ : DenseRange fun x : α × β × γ => ((x.1 : Completion α), ((x.2.1 : Completion β), (x.2.2 : Completion γ))) := - denseRange_coe.prod_map denseRange_coe₂ + denseRange_coe.prodMap denseRange_coe₂ @[elab_as_elim] theorem induction_on {p : Completion α → Prop} (a : Completion α) (hp : IsClosed { a | p a }) diff --git a/test/Continuity.lean b/test/Continuity.lean index 70da467d763ee..bf0d389a95e2a 100644 --- a/test/Continuity.lean +++ b/test/Continuity.lean @@ -53,8 +53,7 @@ example (f : C(X, Y)) (g : C(Y, Z)) : Continuous (g ∘ f) := by continuity example (f : C(X, Y)) (g : C(X, Z)) : Continuous (fun x => (f x, g x)) := by continuity -example (f : C(X, Y)) (g : C(W, Z)) : Continuous (Prod.map f g) := --by continuity - f.continuous.prod_map g.continuous +example (f : C(X, Y)) (g : C(W, Z)) : Continuous (Prod.map f g) := by continuity example (f : ∀ i, C(X, X' i)) : Continuous (fun a i => f i a) := by continuity From 4c163af32071778f5596ac89914e8ae45853dd63 Mon Sep 17 00:00:00 2001 From: Moritz Firsching Date: Mon, 7 Oct 2024 17:59:24 +0000 Subject: [PATCH 37/43] chore(Mathlib/Data/Nat/Defs): remove `Nat.add_def` (#17488) ... since this is a duplicate of `Nat.add_eq`. Drive-by: clean up a proof where add_def was used a bit. Co-authored-by: Moritz Firsching --- Mathlib/Data/Nat/Defs.lean | 3 +-- Mathlib/NumberTheory/Padics/RingHoms.lean | 2 +- 2 files changed, 2 insertions(+), 3 deletions(-) diff --git a/Mathlib/Data/Nat/Defs.lean b/Mathlib/Data/Nat/Defs.lean index d7f4c31154a69..2d6cda97e85fd 100644 --- a/Mathlib/Data/Nat/Defs.lean +++ b/Mathlib/Data/Nat/Defs.lean @@ -223,8 +223,7 @@ attribute [simp] le_add_left le_add_right Nat.lt_add_left_iff_pos Nat.lt_add_rig -- Sometimes a bare `Nat.add` or similar appears as a consequence of unfolding during pattern -- matching. These lemmas package them back up as typeclass mediated operations. --- TODO: This is a duplicate of `Nat.add_eq` -@[simp] lemma add_def : Nat.add m n = m + n := rfl +@[deprecated (since := "2024-04-05")] alias add_def := add_eq -- We want to use these two lemmas earlier than the lemmas simp can prove them with @[simp, nolint simpNF] protected lemma add_eq_left : a + b = a ↔ b = 0 := by omega diff --git a/Mathlib/NumberTheory/Padics/RingHoms.lean b/Mathlib/NumberTheory/Padics/RingHoms.lean index 31ba8f2ff384a..8f65a9cbc788e 100644 --- a/Mathlib/NumberTheory/Padics/RingHoms.lean +++ b/Mathlib/NumberTheory/Padics/RingHoms.lean @@ -256,7 +256,7 @@ theorem toZMod_spec : x - (ZMod.cast (toZMod x) : ℤ_[p]) ∈ maximalIdeal ℤ_ dsimp [toZMod, toZModHom] rcases Nat.exists_eq_add_of_lt hp_prime.1.pos with ⟨p', rfl⟩ change ↑((_ : ZMod (0 + p' + 1)).val) = (_ : ℤ_[0 + p' + 1]) - simp only [ZMod.val_natCast, add_zero, add_def, Nat.cast_inj, zero_add] + rw [Nat.cast_inj] apply mod_eq_of_lt simpa only [zero_add] using zmodRepr_lt_p x From ae8a2a18b7da23a9eae6794db9a1197b2af24972 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Martin=20Dvo=C5=99=C3=A1k?= Date: Mon, 7 Oct 2024 17:59:25 +0000 Subject: [PATCH 38/43] feat(Algebra/Order/BigOperators/Group/*): `max_prod_le`, `prod_min_le` (#17496) `List`, `Multiset`, `Finset` : `max_prod_le`, `max_sum_le`, `prod_min_le`, `prod_sum_le` --- .../Order/BigOperators/Group/Finset.lean | 10 ++++++++++ .../Algebra/Order/BigOperators/Group/List.lean | 18 ++++++++++++++++++ .../Order/BigOperators/Group/Multiset.lean | 14 ++++++++++++++ 3 files changed, 42 insertions(+) diff --git a/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean b/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean index c7a0670ed2b05..d5cdc6b8f86e7 100644 --- a/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean +++ b/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean @@ -210,6 +210,16 @@ theorem prod_le_prod_fiberwise_of_prod_fiber_le_one' {t : Finset ι'} {g : ι end OrderedCommMonoid +@[to_additive] +lemma max_prod_le [LinearOrderedCommMonoid M] {f g : ι → M} {s : Finset ι} : + max (s.prod f) (s.prod g) ≤ s.prod (fun i ↦ max (f i) (g i)) := + Multiset.max_prod_le + +@[to_additive] +lemma prod_min_le [LinearOrderedCommMonoid M] {f g : ι → M} {s : Finset ι} : + s.prod (fun i ↦ min (f i) (g i)) ≤ min (s.prod f) (s.prod g) := + Multiset.prod_min_le + theorem abs_sum_le_sum_abs {G : Type*} [LinearOrderedAddCommGroup G] (f : ι → G) (s : Finset ι) : |∑ i ∈ s, f i| ≤ ∑ i ∈ s, |f i| := le_sum_of_subadditive _ abs_zero abs_add s f diff --git a/Mathlib/Algebra/Order/BigOperators/Group/List.lean b/Mathlib/Algebra/Order/BigOperators/Group/List.lean index 900c60857b3b1..94dd5a115870c 100644 --- a/Mathlib/Algebra/Order/BigOperators/Group/List.lean +++ b/Mathlib/Algebra/Order/BigOperators/Group/List.lean @@ -119,6 +119,24 @@ lemma one_le_prod_of_one_le [Preorder M] [CovariantClass M M (· * ·) (· ≤ rw [prod_cons] exact one_le_mul (hl₁ hd (mem_cons_self hd tl)) (ih fun x h => hl₁ x (mem_cons_of_mem hd h)) +@[to_additive] +lemma max_prod_le (l : List α) (f g : α → M) [LinearOrder M] + [CovariantClass M M (· * ·) (· ≤ ·)] [CovariantClass M M (Function.swap (· * ·)) (· ≤ ·)] : + max (l.map f).prod (l.map g).prod ≤ (l.map fun i ↦ max (f i) (g i)).prod := by + rw [max_le_iff] + constructor <;> apply List.prod_le_prod' <;> intros + · apply le_max_left + · apply le_max_right + +@[to_additive] +lemma prod_min_le [LinearOrder M] [CovariantClass M M (· * ·) (· ≤ ·)] + [CovariantClass M M (Function.swap (· * ·)) (· ≤ ·)] (l : List α) (f g : α → M) : + (l.map fun i ↦ min (f i) (g i)).prod ≤ min (l.map f).prod (l.map g).prod := by + rw [le_min_iff] + constructor <;> apply List.prod_le_prod' <;> intros + · apply min_le_left + · apply min_le_right + end Monoid -- TODO: develop theory of tropical rings diff --git a/Mathlib/Algebra/Order/BigOperators/Group/Multiset.lean b/Mathlib/Algebra/Order/BigOperators/Group/Multiset.lean index 52ae9a3e331ce..ad405fe75fb40 100644 --- a/Mathlib/Algebra/Order/BigOperators/Group/Multiset.lean +++ b/Mathlib/Algebra/Order/BigOperators/Group/Multiset.lean @@ -156,6 +156,20 @@ lemma max_le_of_forall_le {α : Type*} [LinearOrder α] [OrderBot α] (l : Multi induction l using Quotient.inductionOn simpa using List.max_le_of_forall_le _ _ h +@[to_additive] +lemma max_prod_le [LinearOrderedCommMonoid α] {s : Multiset ι} {f g : ι → α} : + max (s.map f).prod (s.map g).prod ≤ (s.map fun i ↦ max (f i) (g i)).prod := by + obtain ⟨l⟩ := s + simp_rw [Multiset.quot_mk_to_coe'', Multiset.map_coe, Multiset.prod_coe] + apply List.max_prod_le + +@[to_additive] +lemma prod_min_le [LinearOrderedCommMonoid α] {s : Multiset ι} {f g : ι → α} : + (s.map fun i ↦ min (f i) (g i)).prod ≤ min (s.map f).prod (s.map g).prod := by + obtain ⟨l⟩ := s + simp_rw [Multiset.quot_mk_to_coe'', Multiset.map_coe, Multiset.prod_coe] + apply List.prod_min_le + lemma abs_sum_le_sum_abs [LinearOrderedAddCommGroup α] {s : Multiset α} : |s.sum| ≤ (s.map abs).sum := le_sum_of_subadditive _ abs_zero abs_add s From ed3c0529273c12f818ddf69e616fa08d3f612e4e Mon Sep 17 00:00:00 2001 From: Dagur Asgeirsson Date: Mon, 7 Oct 2024 17:59:27 +0000 Subject: [PATCH 39/43] chore(Condensed): fix typos in docstrings (#17497) --- Mathlib/Condensed/Discrete/Module.lean | 12 ++++++------ 1 file changed, 6 insertions(+), 6 deletions(-) diff --git a/Mathlib/Condensed/Discrete/Module.lean b/Mathlib/Condensed/Discrete/Module.lean index 0c30c3516af79..3bcf129729c30 100644 --- a/Mathlib/Condensed/Discrete/Module.lean +++ b/Mathlib/Condensed/Discrete/Module.lean @@ -76,13 +76,13 @@ abbrev functor : ModuleCat R ⥤ CondensedMod.{u} R := CompHausLike.LocallyConstantModule.functor.{u+1, u} R (fun _ _ _ ↦ ((CompHaus.effectiveEpi_tfae _).out 0 2).mp) -/-- Auxilary definition for `functorIsoDiscrete`. -/ +/-- Auxiliary definition for `functorIsoDiscrete`. -/ noncomputable def functorIsoDiscreteAux₁ (M : ModuleCat.{u+1} R) : M ≅ (ModuleCat.of R (LocallyConstant (CompHaus.of PUnit.{u+1}) M)) where hom := constₗ R inv := evalₗ R PUnit.unit -/-- Auxilary definition for `functorIsoDiscrete`. -/ +/-- Auxiliary definition for `functorIsoDiscrete`. -/ noncomputable def functorIsoDiscreteAux₂ (M : ModuleCat R) : (discrete _).obj M ≅ (discrete _).obj (ModuleCat.of R (LocallyConstant (CompHaus.of PUnit.{u+1}) M)) := @@ -103,7 +103,7 @@ instance (M : ModuleCat R) : IsIso ((forget R).map rw [essImage_eq_of_natIso CondensedSet.LocallyConstant.iso.symm] exact obj_mem_essImage CondensedSet.LocallyConstant.functor M -/-- Auxilary definition for `functorIsoDiscrete`. -/ +/-- Auxiliary definition for `functorIsoDiscrete`. -/ noncomputable def functorIsoDiscreteComponents (M : ModuleCat R) : (discrete _).obj M ≅ (functor R).obj M := have : (Condensed.forget R).ReflectsIsomorphisms := @@ -183,13 +183,13 @@ abbrev functor : ModuleCat R ⥤ LightCondMod.{u} R := CompHausLike.LocallyConstantModule.functor.{u, u} R (fun _ _ _ ↦ (LightProfinite.effectiveEpi_iff_surjective _).mp) -/-- Auxilary definition for `functorIsoDiscrete`. -/ +/-- Auxiliary definition for `functorIsoDiscrete`. -/ noncomputable def functorIsoDiscreteAux₁ (M : ModuleCat.{u} R) : M ≅ (ModuleCat.of R (LocallyConstant (LightProfinite.of PUnit.{u+1}) M)) where hom := constₗ R inv := evalₗ R PUnit.unit -/-- Auxilary definition for `functorIsoDiscrete`. -/ +/-- Auxiliary definition for `functorIsoDiscrete`. -/ noncomputable def functorIsoDiscreteAux₂ (M : ModuleCat.{u} R) : (discrete _).obj M ≅ (discrete _).obj (ModuleCat.of R (LocallyConstant (LightProfinite.of PUnit.{u+1}) M)) := @@ -216,7 +216,7 @@ instance (M : ModuleCat R) : rw [essImage_eq_of_natIso LightCondSet.LocallyConstant.iso.symm] exact obj_mem_essImage LightCondSet.LocallyConstant.functor M -/-- Auxilary definition for `functorIsoDiscrete`. -/ +/-- Auxiliary definition for `functorIsoDiscrete`. -/ noncomputable def functorIsoDiscreteComponents (M : ModuleCat R) : (discrete _).obj M ≅ (functor R).obj M := have : (LightCondensed.forget R).ReflectsIsomorphisms := From 4c46b674f8a477e2ea3580135ed4c51a2e620f91 Mon Sep 17 00:00:00 2001 From: grunweg Date: Mon, 7 Oct 2024 17:59:28 +0000 Subject: [PATCH 40/43] chore(RamificationInertia): remove superfluous DecidableEq (#17501) Found by the linter in #10235. --- Mathlib/NumberTheory/RamificationInertia.lean | 11 ++++++----- 1 file changed, 6 insertions(+), 5 deletions(-) diff --git a/Mathlib/NumberTheory/RamificationInertia.lean b/Mathlib/NumberTheory/RamificationInertia.lean index 738d4accca7ef..670a2c3ba24a8 100644 --- a/Mathlib/NumberTheory/RamificationInertia.lean +++ b/Mathlib/NumberTheory/RamificationInertia.lean @@ -797,11 +797,11 @@ section tower variable {R S T : Type*} [CommRing R] [CommRing S] [CommRing T] -theorem ramificationIdx_tower [IsDedekindDomain S] [DecidableEq (Ideal S)] - [IsDedekindDomain T] [DecidableEq (Ideal T)]{f : R →+* S} {g : S →+* T} +theorem ramificationIdx_tower [IsDedekindDomain S] [IsDedekindDomain T] {f : R →+* S} {g : S →+* T} {p : Ideal R} {P : Ideal S} {Q : Ideal T} [hpm : P.IsPrime] [hqm : Q.IsPrime] (hg0 : map g P ≠ ⊥) (hfg : map (g.comp f) p ≠ ⊥) (hg : map g P ≤ Q) : ramificationIdx (g.comp f) p Q = ramificationIdx f p P * ramificationIdx g P Q := by + classical have hf0 : map f p ≠ ⊥ := ne_bot_of_map_ne_bot (Eq.mp (congrArg (fun I ↦ I ≠ ⊥) (map_map f g).symm) hfg) have hp0 : P ≠ ⊥ := ne_bot_of_map_ne_bot hg0 @@ -836,12 +836,13 @@ variable [Algebra R S] [Algebra S T] [Algebra R T] [IsScalarTower R S T] /-- Let `T / S / R` be a tower of algebras, `p, P, Q` be ideals in `R, S, T` respectively, and `P` and `Q` are prime. If `P = Q ∩ S`, then `e (Q | p) = e (P | p) * e (Q | P)`. -/ -theorem ramificationIdx_algebra_tower [IsDedekindDomain S] [DecidableEq (Ideal S)] - [IsDedekindDomain T] [DecidableEq (Ideal T)] {p : Ideal R} {P : Ideal S} {Q : Ideal T} - [hpm : P.IsPrime] [hqm : Q.IsPrime] (hg0 : map (algebraMap S T) P ≠ ⊥) +theorem ramificationIdx_algebra_tower [IsDedekindDomain S] [IsDedekindDomain T] + {p : Ideal R} {P : Ideal S} {Q : Ideal T} [hpm : P.IsPrime] [hqm : Q.IsPrime] + (hg0 : map (algebraMap S T) P ≠ ⊥) (hfg : map (algebraMap R T) p ≠ ⊥) (hg : map (algebraMap S T) P ≤ Q) : ramificationIdx (algebraMap R T) p Q = ramificationIdx (algebraMap R S) p P * ramificationIdx (algebraMap S T) P Q := by + classical rw [IsScalarTower.algebraMap_eq R S T] at hfg ⊢ exact ramificationIdx_tower hg0 hfg hg From 3ce877a139a51fb84abba4cd2d86078d64e848f9 Mon Sep 17 00:00:00 2001 From: grunweg Date: Mon, 7 Oct 2024 17:59:29 +0000 Subject: [PATCH 41/43] chore(DomMulAct): remove decidability assumptions (#17510) Found by the linter in #10235. --- Mathlib/GroupTheory/Perm/DomMulAct.lean | 7 +++++-- 1 file changed, 5 insertions(+), 2 deletions(-) diff --git a/Mathlib/GroupTheory/Perm/DomMulAct.lean b/Mathlib/GroupTheory/Perm/DomMulAct.lean index 34ebef212fe7d..4bfd4c758ab0d 100644 --- a/Mathlib/GroupTheory/Perm/DomMulAct.lean +++ b/Mathlib/GroupTheory/Perm/DomMulAct.lean @@ -88,14 +88,14 @@ lemma stabilizerMulEquiv_apply (g : (stabilizer (Perm α)ᵈᵐᵃ f)ᵐᵒᵖ) section Fintype -variable [Fintype α] [DecidableEq α] [DecidableEq ι] +variable [Fintype α] open Nat variable (f) /-- The cardinality of the type of permutations preserving a function -/ -theorem stabilizer_card [Fintype ι] : +theorem stabilizer_card [DecidableEq α] [DecidableEq ι] [Fintype ι] : Fintype.card {g : Perm α // f ∘ g = f} = ∏ i, (Fintype.card {a // f a = i})! := by -- rewriting via Nat.card because Fintype instance is not found rw [← Nat.card_eq_fintype_card, @@ -108,9 +108,12 @@ theorem stabilizer_card [Fintype ι] : /-- The cardinality of the set of permutations preserving a function -/ theorem stabilizer_ncard [Fintype ι] : Set.ncard {g : Perm α | f ∘ g = f} = ∏ i, (Set.ncard {a | f a = i})! := by + classical simp only [← Set.Nat.card_coe_set_eq, Set.coe_setOf, card_eq_fintype_card] exact stabilizer_card f +variable [DecidableEq α] [DecidableEq ι] + /-- The cardinality of the type of permutations preserving a function (without the finiteness assumption on target)-/ theorem stabilizer_card': From 312b6226d3fedc5fe0b06eb81fcf0476ff514c5a Mon Sep 17 00:00:00 2001 From: Jiang Jiedong <107380768+jjdishere@users.noreply.github.com> Date: Mon, 7 Oct 2024 20:10:01 +0000 Subject: [PATCH 42/43] feat(FieldTheory/Minpoly): add theorem for minpoly of -x (#17367) Add theorem `minpoly A (- x) = (-1) ^ (natDegree (minpoly A x)) * (minpoly A x).comp (- X)` Co-authored-by: jjdishere --- Mathlib/Algebra/Polynomial/RingDivision.lean | 20 ++++++++++++++++++ Mathlib/FieldTheory/Minpoly/Field.lean | 21 +++++++++++++++++++ .../IntegralClosure/Algebra/Basic.lean | 14 +++++++++++++ 3 files changed, 55 insertions(+) diff --git a/Mathlib/Algebra/Polynomial/RingDivision.lean b/Mathlib/Algebra/Polynomial/RingDivision.lean index ab1ca1c76d94d..87a240f259415 100644 --- a/Mathlib/Algebra/Polynomial/RingDivision.lean +++ b/Mathlib/Algebra/Polynomial/RingDivision.lean @@ -548,6 +548,26 @@ theorem rootMultiplicity_mul' {p q : R[X]} {x : R} theorem Monic.comp_X_sub_C {p : R[X]} (hp : p.Monic) (r : R) : (p.comp (X - C r)).Monic := by simpa using hp.comp_X_add_C (-r) +@[simp] +theorem comp_neg_X_leadingCoeff_eq (p : R[X]) : + (p.comp (-X)).leadingCoeff = (-1) ^ p.natDegree * p.leadingCoeff := by + nontriviality R + by_cases h : p = 0 + · simp [h] + rw [Polynomial.leadingCoeff, natDegree_comp_eq_of_mul_ne_zero, coeff_comp_degree_mul_degree] <;> + simp [mul_comm, h] + +theorem Monic.neg_one_pow_natDegree_mul_comp_neg_X {p : R[X]} (hp : p.Monic) : + ((-1) ^ p.natDegree * p.comp (-X)).Monic := by + simp only [Monic] + calc + ((-1) ^ p.natDegree * p.comp (-X)).leadingCoeff = + (p.comp (-X) * C ((-1) ^ p.natDegree)).leadingCoeff := by + simp [mul_comm] + _ = 1 := by + apply monic_mul_C_of_leadingCoeff_mul_eq_one + simp [← pow_add, hp] + variable [IsDomain R] {p q : R[X]} @[simp] diff --git a/Mathlib/FieldTheory/Minpoly/Field.lean b/Mathlib/FieldTheory/Minpoly/Field.lean index 3a33353b9af98..b129b8a0ecbc5 100644 --- a/Mathlib/FieldTheory/Minpoly/Field.lean +++ b/Mathlib/FieldTheory/Minpoly/Field.lean @@ -148,6 +148,27 @@ theorem sub_algebraMap {B : Type*} [CommRing B] [Algebra A B] {x : B} (a : A) : minpoly A (x - algebraMap A B a) = (minpoly A x).comp (X + C a) := by simpa [sub_eq_add_neg] using add_algebraMap x (-a) +theorem neg {B : Type*} [CommRing B] [Algebra A B] (x : B) : + minpoly A (- x) = (-1) ^ (natDegree (minpoly A x)) * (minpoly A x).comp (- X) := by + by_cases hx : IsIntegral A x + · refine (minpoly.unique _ _ ((minpoly.monic hx).neg_one_pow_natDegree_mul_comp_neg_X) + ?_ fun q qmo hq => ?_).symm + · simp [aeval_comp] + · have : (Polynomial.aeval x) ((-1) ^ q.natDegree * q.comp (- X)) = 0 := by + simpa [aeval_comp] using hq + have H := minpoly.min A x qmo.neg_one_pow_natDegree_mul_comp_neg_X this + have n1 := ((minpoly.monic hx).neg_one_pow_natDegree_mul_comp_neg_X).ne_zero + have n2 := qmo.neg_one_pow_natDegree_mul_comp_neg_X.ne_zero + rw [degree_eq_natDegree qmo.ne_zero, + degree_eq_natDegree n1, natDegree_mul (by simp) (right_ne_zero_of_mul n1), natDegree_comp] + rw [degree_eq_natDegree (minpoly.ne_zero hx), + degree_eq_natDegree qmo.neg_one_pow_natDegree_mul_comp_neg_X.ne_zero, + natDegree_mul (by simp) (right_ne_zero_of_mul n2), natDegree_comp] at H + simpa using H + · rw [minpoly.eq_zero hx, minpoly.eq_zero, zero_comp] + · simp only [natDegree_zero, pow_zero, mul_zero] + · exact IsIntegral.neg_iff.not.mpr hx + section AlgHomFintype /-- A technical finiteness result. -/ diff --git a/Mathlib/RingTheory/IntegralClosure/Algebra/Basic.lean b/Mathlib/RingTheory/IntegralClosure/Algebra/Basic.lean index ac6ad06cc3419..37343ddf2c611 100644 --- a/Mathlib/RingTheory/IntegralClosure/Algebra/Basic.lean +++ b/Mathlib/RingTheory/IntegralClosure/Algebra/Basic.lean @@ -148,9 +148,23 @@ variable (f : R →+* S) theorem RingHom.IsIntegralElem.neg {x : S} (hx : f.IsIntegralElem x) : f.IsIntegralElem (-x) := hx.of_mem_closure f hx (Subring.neg_mem _ (Subring.subset_closure (Or.inl rfl))) +theorem RingHom.IsIntegralElem.of_neg {x : S} (h : f.IsIntegralElem (-x)) : f.IsIntegralElem x := + neg_neg x ▸ h.neg + +@[simp] +theorem RingHom.IsIntegralElem.neg_iff {x : S} : f.IsIntegralElem (-x) ↔ f.IsIntegralElem x := + ⟨fun h => h.of_neg, fun h => h.neg⟩ + theorem IsIntegral.neg {x : B} (hx : IsIntegral R x) : IsIntegral R (-x) := .of_mem_of_fg _ hx.fg_adjoin_singleton _ (Subalgebra.neg_mem _ <| Algebra.subset_adjoin rfl) +theorem IsIntegral.of_neg {x : B} (hx : IsIntegral R (-x)) : IsIntegral R x := + neg_neg x ▸ hx.neg + +@[simp] +theorem IsIntegral.neg_iff {x : B} : IsIntegral R (-x) ↔ IsIntegral R x := + ⟨IsIntegral.of_neg, IsIntegral.neg⟩ + theorem RingHom.IsIntegralElem.sub {x y : S} (hx : f.IsIntegralElem x) (hy : f.IsIntegralElem y) : f.IsIntegralElem (x - y) := by simpa only [sub_eq_add_neg] using hx.add f (hy.neg f) From a239b9573c31a8a440e726e42bdc39c956ef5e0e Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 7 Oct 2024 20:10:02 +0000 Subject: [PATCH 43/43] chore(*): drop some `[Decidable*]` assumptions (#17396) Related to #10235. --- Mathlib/Algebra/MvPolynomial/PDeriv.lean | 17 +++++++++-------- Mathlib/NumberTheory/JacobiSum/Basic.lean | 5 +++-- 2 files changed, 12 insertions(+), 10 deletions(-) diff --git a/Mathlib/Algebra/MvPolynomial/PDeriv.lean b/Mathlib/Algebra/MvPolynomial/PDeriv.lean index b7aa67a2aad48..2f36da703dccd 100644 --- a/Mathlib/Algebra/MvPolynomial/PDeriv.lean +++ b/Mathlib/Algebra/MvPolynomial/PDeriv.lean @@ -65,12 +65,12 @@ theorem pderiv_def [DecidableEq σ] (i : σ) : pderiv i = mkDerivation R (Pi.sin theorem pderiv_monomial {i : σ} : pderiv i (monomial s a) = monomial (s - single i 1) (a * s i) := by classical - simp only [pderiv_def, mkDerivation_monomial, Finsupp.smul_sum, smul_eq_mul, ← smul_mul_assoc, - ← (monomial _).map_smul] - refine (Finset.sum_eq_single i (fun j _ hne => ?_) fun hi => ?_).trans ?_ - · simp [Pi.single_eq_of_ne hne] - · rw [Finsupp.not_mem_support_iff] at hi; simp [hi] - · simp + simp only [pderiv_def, mkDerivation_monomial, Finsupp.smul_sum, smul_eq_mul, ← smul_mul_assoc, + ← (monomial _).map_smul] + refine (Finset.sum_eq_single i (fun j _ hne => ?_) fun hi => ?_).trans ?_ + · simp [Pi.single_eq_of_ne hne] + · rw [Finsupp.not_mem_support_iff] at hi; simp [hi] + · simp theorem pderiv_C {i : σ} : pderiv i (C a) = 0 := derivation_C _ _ @@ -115,9 +115,10 @@ theorem pderiv_map {S} [CommSemiring S] {φ : R →+* S} {f : MvPolynomial σ R} · simp [eq] · simp [eq, h] -lemma pderiv_rename {τ : Type*} [DecidableEq τ] [DecidableEq σ] {f : σ → τ} - (hf : Function.Injective f) (x : σ) (p : MvPolynomial σ R) : +lemma pderiv_rename {τ : Type*} {f : σ → τ} (hf : Function.Injective f) + (x : σ) (p : MvPolynomial σ R) : pderiv (f x) (rename f p) = rename f (pderiv x p) := by + classical induction' p using MvPolynomial.induction_on with a p q hp hq p a h · simp · simp [hp, hq] diff --git a/Mathlib/NumberTheory/JacobiSum/Basic.lean b/Mathlib/NumberTheory/JacobiSum/Basic.lean index f9a173578aa98..9088bfbe869f8 100644 --- a/Mathlib/NumberTheory/JacobiSum/Basic.lean +++ b/Mathlib/NumberTheory/JacobiSum/Basic.lean @@ -270,7 +270,7 @@ lemma MulChar.exists_apply_sub_one_mul_apply_sub_one {n : ℕ} (hn : n ≠ 0) { with values in an integral domain `R` and `μ` is a primitive `n`th root of unity in `R`, then `J(χ,ψ) = -1 + z*(μ - 1)^2` for some `z ∈ ℤ[μ] ⊆ R`. (We assume that `#F ≡ 1 mod n`.) Note that we do not state this as a divisbility in `R`, as this would give a weaker statement. -/ -lemma exists_jacobiSum_eq_neg_one_add [DecidableEq F] {n : ℕ} (hn : 2 < n) {χ ψ : MulChar F R} +lemma exists_jacobiSum_eq_neg_one_add {n : ℕ} (hn : 2 < n) {χ ψ : MulChar F R} {μ : R} (hχ : χ ^ n = 1) (hψ : ψ ^ n = 1) (hn' : n ∣ Fintype.card F - 1) (hμ : IsPrimitiveRoot μ n) : ∃ z ∈ Algebra.adjoin ℤ {μ}, jacobiSum χ ψ = -1 + z * (μ - 1) ^ 2 := by @@ -286,7 +286,8 @@ lemma exists_jacobiSum_eq_neg_one_add [DecidableEq F] {n : ℕ} (hn : 2 < n) {χ rw [hχ₀, jacobiSum_one_nontrivial hψ₀, zero_mul, add_zero] · refine ⟨0, Subalgebra.zero_mem _, ?_⟩ rw [jacobiSum_comm, hψ₀, jacobiSum_one_nontrivial hχ₀, zero_mul, add_zero] - · rw [jacobiSum_eq_aux, MulChar.sum_eq_zero_of_ne_one hχ₀, MulChar.sum_eq_zero_of_ne_one hψ₀, hq] + · classical + rw [jacobiSum_eq_aux, MulChar.sum_eq_zero_of_ne_one hχ₀, MulChar.sum_eq_zero_of_ne_one hψ₀, hq] have H := MulChar.exists_apply_sub_one_mul_apply_sub_one (by omega) hχ hψ hμ have Hcs x := (H x).choose_spec refine ⟨-q * z₁ + ∑ x ∈ (univ \ {0, 1} : Finset F), (H x).choose, ?_, ?_⟩