From 7a8cf3c4b324fbfaf375dccf41c5be3c3099c3bb Mon Sep 17 00:00:00 2001 From: Christian Merten Date: Wed, 16 Oct 2024 05:50:11 +0000 Subject: [PATCH 001/505] feat(AlgebraicGeometry/RingHomProperties): a property induced from `Locally Q` is equivalent to locally `Q` (#17305) If `P` is a morphism property of schemes associated to the property `Locally Q` for a property of ring homomorphisms `Q`, then for a morphism `f`, `P f` is equivalent to: for every point `x` in the source there exist open affine neighborhoods around `x` and `f.val.base x` respectively such that the induced map of rings satisfies `Q`. This guarantees that we may define properties of scheme morphisms explicitly as described above for properties of ring homomorphisms that are not necessarily local at the target and still use the API for `HasRingHomProperty`. --- .../Morphisms/RingHomProperties.lean | 57 ++++++++++++++++--- 1 file changed, 48 insertions(+), 9 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean b/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean index 9a5b418ed0cd6..93a2ff38a6b0a 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean @@ -5,6 +5,7 @@ Authors: Andrew Yang -/ import Mathlib.AlgebraicGeometry.Morphisms.Basic import Mathlib.RingTheory.LocalProperties.Basic +import Mathlib.RingTheory.RingHom.Locally /-! @@ -182,6 +183,17 @@ namespace HasRingHomProperty variable (P : MorphismProperty Scheme.{u}) {Q} [HasRingHomProperty P Q] variable {X Y Z : Scheme.{u}} (f : X ⟶ Y) (g : Y ⟶ Z) +lemma copy {P' : MorphismProperty Scheme.{u}} + {Q' : ∀ {R S : Type u} [CommRing R] [CommRing S], (R →+* S) → Prop} + (e : P = P') (e' : ∀ {R S : Type u} [CommRing R] [CommRing S] (f : R →+* S), Q f ↔ Q' f) : + HasRingHomProperty P' Q' := by + subst e + have heq : @Q = @Q' := by + ext R S _ _ f + exact (e' f) + rw [← heq] + infer_instance + lemma eq_affineLocally : P = affineLocally Q := eq_affineLocally' @[local instance] @@ -416,25 +428,52 @@ instance respects_isOpenImmersion : P.Respects @IsOpenImmersion where rw [show f ≫ i = f ≫ e.hom ≫ i.opensRange.ι by simp [e], ← Category.assoc] exact respects_isOpenImmersion_aux _ (by rwa [P.cancel_right_of_respectsIso]) -open RingHom in -/-- `P` can be checked locally around points of the source. -/ -lemma iff_exists_appLE : P f ↔ - ∀ (x : X), ∃ (U : Y.affineOpens) (V : X.affineOpens) (_ : x ∈ V.1) (e : V.1 ≤ f ⁻¹ᵁ U.1), +open RingHom + +omit [HasRingHomProperty P Q] in +/-- If `P` is induced by `Locally Q`, it suffices to check `Q` on affine open sets locally around +points of the source. -/ +lemma iff_exists_appLE_locally (hQi : RespectsIso Q) [HasRingHomProperty P (Locally Q)] : + P f ↔ ∀ (x : X), ∃ (U : Y.affineOpens) (V : X.affineOpens) (_ : x ∈ V.1) (e : V.1 ≤ f ⁻¹ᵁ U.1), Q (f.appLE U V e) := by refine ⟨fun hf x ↦ ?_, fun hf ↦ (IsLocalAtSource.iff_exists_resLE (P := P)).mpr <| fun x ↦ ?_⟩ · obtain ⟨U, hU, hfx, _⟩ := Opens.isBasis_iff_nbhd.mp (isBasis_affine_open Y) (Opens.mem_top <| f.val.base x) obtain ⟨V, hV, hx, e⟩ := Opens.isBasis_iff_nbhd.mp (isBasis_affine_open X) (show x ∈ f ⁻¹ᵁ U from hfx) - use ⟨U, hU⟩, ⟨V, hV⟩, hx, e - apply appLE P f hf + simp_rw [HasRingHomProperty.iff_appLE (P := P), locally_iff_isLocalization hQi] at hf + obtain ⟨s, hs, hfs⟩ := hf ⟨U, hU⟩ ⟨V, hV⟩ e + apply iSup_basicOpen_of_span_eq_top at hs + have : x ∈ (⨆ i ∈ s, X.basicOpen i) := hs.symm ▸ hx + have : ∃ r ∈ s, x ∈ X.basicOpen r := by simpa using this + obtain ⟨r, hr, hrs⟩ := this + refine ⟨⟨U, hU⟩, ⟨X.basicOpen r, hV.basicOpen r⟩, hrs, (X.basicOpen_le r).trans e, ?_⟩ + rw [← f.appLE_map e (homOfLE (X.basicOpen_le r)).op] + haveI : IsLocalization.Away r Γ(X, X.basicOpen r) := hV.isLocalization_basicOpen r + exact hfs r hr _ · obtain ⟨U, V, hxV, e, hf⟩ := hf x use U, V, hxV, e simp only [iff_of_isAffine (P := P), Scheme.Hom.appLE, homOfLE_leOfHom] at hf ⊢ - haveI : (toMorphismProperty Q).RespectsIso := toMorphismProperty_respectsIso_iff.mp <| + haveI : (toMorphismProperty (Locally Q)).RespectsIso := toMorphismProperty_respectsIso_iff.mp <| (isLocal_ringHomProperty P).respectsIso - exact (MorphismProperty.arrow_mk_iso_iff (toMorphismProperty Q) (arrowResLEAppIso f U V e)).mpr - hf + exact (MorphismProperty.arrow_mk_iso_iff (toMorphismProperty (Locally Q)) + (arrowResLEAppIso f U V e)).mpr (locally_of hQi _ hf) + +/-- `P` can be checked locally around points of the source. -/ +lemma iff_exists_appLE : P f ↔ + ∀ (x : X), ∃ (U : Y.affineOpens) (V : X.affineOpens) (_ : x ∈ V.1) (e : V.1 ≤ f ⁻¹ᵁ U.1), + Q (f.appLE U V e) := by + haveI inst : HasRingHomProperty P Q := inferInstance + haveI : HasRingHomProperty P (Locally Q) := by + apply @copy (P' := P) (Q := Q) (Q' := Locally Q) + · infer_instance + · rfl + · intro R S _ _ f + exact (locally_iff_of_localizationSpanTarget (isLocal_ringHomProperty P).respectsIso + (isLocal_ringHomProperty P).OfLocalizationSpanTarget _).symm + rw [iff_exists_appLE_locally (P := P) ] + haveI : HasRingHomProperty P Q := inst + apply (isLocal_ringHomProperty P (Q := Q)).respectsIso end HasRingHomProperty From b3db14fd6ff41d064b5c86d29fb7749d24a3347b Mon Sep 17 00:00:00 2001 From: Christian Merten Date: Wed, 16 Oct 2024 05:50:13 +0000 Subject: [PATCH 002/505] feat(CategoryTheory/Sites): infima of pretopologies (#17734) Defines a complete lattice structure on pretopologies with good definitional properties for bottom, top and infima. --- .../CategoryTheory/Sites/Grothendieck.lean | 5 ++ Mathlib/CategoryTheory/Sites/Pretopology.lean | 66 ++++++++++++++++++- 2 files changed, 69 insertions(+), 2 deletions(-) diff --git a/Mathlib/CategoryTheory/Sites/Grothendieck.lean b/Mathlib/CategoryTheory/Sites/Grothendieck.lean index aaf3d4a203a5b..c29e44a6f4bf3 100644 --- a/Mathlib/CategoryTheory/Sites/Grothendieck.lean +++ b/Mathlib/CategoryTheory/Sites/Grothendieck.lean @@ -260,6 +260,11 @@ instance : InfSet (GrothendieckTopology C) where apply J.transitive (hS _ ⟨⟨_, _, hJ, rfl⟩, rfl⟩) _ fun Y f hf => h hf _ ⟨⟨_, _, hJ, rfl⟩, rfl⟩ } +lemma mem_sInf (s : Set (GrothendieckTopology C)) {X : C} (S : Sieve X) : + S ∈ sInf s X ↔ ∀ t ∈ s, S ∈ t X := by + show S ∈ sInf (sieves '' s) X ↔ _ + simp + /-- See -/ theorem isGLB_sInf (s : Set (GrothendieckTopology C)) : IsGLB s (sInf s) := by refine @IsGLB.of_image _ _ _ _ sieves ?_ _ _ ?_ diff --git a/Mathlib/CategoryTheory/Sites/Pretopology.lean b/Mathlib/CategoryTheory/Sites/Pretopology.lean index a70e13cc0de8b..85b5f87c3c0a8 100644 --- a/Mathlib/CategoryTheory/Sites/Pretopology.lean +++ b/Mathlib/CategoryTheory/Sites/Pretopology.lean @@ -87,7 +87,7 @@ instance : PartialOrder (Pretopology C) := le_trans := fun _ _ _ h₁₂ h₂₃ => le_def.mpr (le_trans h₁₂ h₂₃) le_antisymm := fun _ _ h₁₂ h₂₁ => Pretopology.ext (le_antisymm h₁₂ h₂₁) } -instance : OrderTop (Pretopology C) where +instance orderTop : OrderTop (Pretopology C) where top := { coverings := fun _ => Set.univ has_isos := fun _ _ _ _ => Set.mem_univ _ @@ -156,6 +156,10 @@ def gi : GaloisInsertion (toGrothendieck C) (ofGrothendieck C) where choice x _ := toGrothendieck C x choice_eq _ _ := rfl +lemma mem_ofGrothendieck (t : GrothendieckTopology C) {X : C} (S : Presieve X) : + S ∈ ofGrothendieck C t X ↔ Sieve.generate S ∈ t X := + Iff.rfl + /-- The trivial pretopology, in which the coverings are exactly singleton isomorphisms. This topology is also known as the indiscrete, coarse, or chaotic topology. @@ -194,7 +198,7 @@ def trivial : Pretopology C where rw [hTi] apply singleton.mk -instance : OrderBot (Pretopology C) where +instance orderBot : OrderBot (Pretopology C) where bot := trivial C bot_le K X R := by rintro ⟨Y, f, hf, rfl⟩ @@ -204,6 +208,64 @@ instance : OrderBot (Pretopology C) where theorem toGrothendieck_bot : toGrothendieck C ⊥ = ⊥ := (gi C).gc.l_bot +instance : InfSet (Pretopology C) where + sInf T := { + coverings := sInf (coverings '' T) + has_isos := fun X Y f _ ↦ by + simp only [sInf_apply, Set.iInf_eq_iInter, Set.iInter_coe_set, Set.mem_image, + Set.iInter_exists, + Set.biInter_and', Set.iInter_iInter_eq_right, Set.mem_iInter] + intro t _ + exact t.has_isos f + pullbacks := fun X Y f S hS ↦ by + simp only [sInf_apply, Set.iInf_eq_iInter, Set.iInter_coe_set, Set.mem_image, + Set.iInter_exists, Set.biInter_and', Set.iInter_iInter_eq_right, Set.mem_iInter] at hS ⊢ + intro t ht + exact t.pullbacks f S (hS t ht) + transitive := fun X S Ti hS hTi ↦ by + simp only [sInf_apply, Set.iInf_eq_iInter, Set.iInter_coe_set, Set.mem_image, + Set.iInter_exists, Set.biInter_and', Set.iInter_iInter_eq_right, Set.mem_iInter] at hS hTi ⊢ + intro t ht + exact t.transitive S Ti (hS t ht) (fun Y f H ↦ hTi f H t ht) + } + +lemma mem_sInf (T : Set (Pretopology C)) {X : C} (S : Presieve X) : + S ∈ sInf T X ↔ ∀ t ∈ T, S ∈ t X := by + show S ∈ sInf (Pretopology.coverings '' T) X ↔ _ + simp + +lemma sInf_ofGrothendieck (T : Set (GrothendieckTopology C)) : + ofGrothendieck C (sInf T) = sInf (ofGrothendieck C '' T) := by + ext X S + simp [mem_sInf, mem_ofGrothendieck, GrothendieckTopology.mem_sInf] + +lemma isGLB_sInf (T : Set (Pretopology C)) : IsGLB T (sInf T) := + IsGLB.of_image (f := coverings) Iff.rfl (_root_.isGLB_sInf _) + +/-- The complete lattice structure on pretopologies. This is induced by the `InfSet` instance, but +with good definitional equalities for `⊥`, `⊤` and `⊓`. -/ +instance : CompleteLattice (Pretopology C) where + __ := orderBot C + __ := orderTop C + inf t₁ t₂ := { + coverings := fun X ↦ t₁.coverings X ∩ t₂.coverings X + has_isos := fun _ _ f _ ↦ + ⟨t₁.has_isos f, t₂.has_isos f⟩ + pullbacks := fun _ _ f S hS ↦ + ⟨t₁.pullbacks f S hS.left, t₂.pullbacks f S hS.right⟩ + transitive := fun _ S Ti hS hTi ↦ + ⟨t₁.transitive S Ti hS.left (fun _ f H ↦ (hTi f H).left), + t₂.transitive S Ti hS.right (fun _ f H ↦ (hTi f H).right)⟩ + } + inf_le_left _ _ _ _ hS := hS.left + inf_le_right _ _ _ _ hS := hS.right + le_inf _ _ _ hts htr X _ hS := ⟨hts X hS, htr X hS⟩ + __ := completeLatticeOfInf _ (isGLB_sInf C) + +lemma mem_inf (t₁ t₂ : Pretopology C) {X : C} (S : Presieve X) : + S ∈ (t₁ ⊓ t₂) X ↔ S ∈ t₁ X ∧ S ∈ t₂ X := + Iff.rfl + end Pretopology end CategoryTheory From a5bdd27269ae4202c997b47d7879fc90089abbe6 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Wed, 16 Oct 2024 05:50:14 +0000 Subject: [PATCH 003/505] chore: deprecate Fin.ofNat'' (#17759) Now that `NeZero` is in Lean, `Fin.ofNat'` already means the right thing. --- Mathlib/Algebra/Group/Fin/Basic.lean | 2 +- Mathlib/Data/Fin/Basic.lean | 15 +++------------ 2 files changed, 4 insertions(+), 13 deletions(-) diff --git a/Mathlib/Algebra/Group/Fin/Basic.lean b/Mathlib/Algebra/Group/Fin/Basic.lean index 9385f63b70721..06f0bb2f7ad67 100644 --- a/Mathlib/Algebra/Group/Fin/Basic.lean +++ b/Mathlib/Algebra/Group/Fin/Basic.lean @@ -43,7 +43,7 @@ instance addCommMonoid (n : ℕ) [NeZero n] : AddCommMonoid (Fin n) where instance instAddMonoidWithOne (n) [NeZero n] : AddMonoidWithOne (Fin n) where __ := inferInstanceAs (AddCommMonoid (Fin n)) - natCast n := Fin.ofNat'' n + natCast i := Fin.ofNat' n i natCast_zero := rfl natCast_succ _ := Fin.ext (add_mod _ _ _) diff --git a/Mathlib/Data/Fin/Basic.lean b/Mathlib/Data/Fin/Basic.lean index edc71f8ff4d06..0da69a90e9ab7 100644 --- a/Mathlib/Data/Fin/Basic.lean +++ b/Mathlib/Data/Fin/Basic.lean @@ -242,14 +242,12 @@ instance {n : ℕ} : WellFoundedRelation (Fin n) := measure (val : Fin n → ℕ) /-- Given a positive `n`, `Fin.ofNat' i` is `i % n` as an element of `Fin n`. -/ +@[deprecated Fin.ofNat' (since := "2024-10-15")] def ofNat'' [NeZero n] (i : ℕ) : Fin n := ⟨i % n, mod_lt _ n.pos_of_neZero⟩ -- Porting note: `Fin.ofNat'` conflicts with something in core (there the hypothesis is `n > 0`), -- so for now we make this double-prime `''`. This is also the reason for the dubious translation. -instance {n : ℕ} [NeZero n] : Zero (Fin n) := ⟨ofNat'' 0⟩ -instance {n : ℕ} [NeZero n] : One (Fin n) := ⟨ofNat'' 1⟩ - /-- The `Fin.val_zero` in `Lean` only applies in `Fin (n+1)`. This one instead uses a `NeZero n` typeclass hypothesis. @@ -422,15 +420,8 @@ instance inhabitedFinOneAdd (n : ℕ) : Inhabited (Fin (1 + n)) := theorem default_eq_zero (n : ℕ) [NeZero n] : (default : Fin n) = 0 := rfl -section from_ad_hoc - -@[simp] lemma ofNat'_zero [NeZero n] : (Fin.ofNat' n 0) = 0 := rfl -@[simp] lemma ofNat'_one [NeZero n] : (Fin.ofNat' n 1) = 1 := rfl - -end from_ad_hoc - instance instNatCast [NeZero n] : NatCast (Fin n) where - natCast n := Fin.ofNat'' n + natCast i := Fin.ofNat' n i lemma natCast_def [NeZero n] (a : ℕ) : (a : Fin n) = ⟨a % n, mod_lt _ n.pos_of_neZero⟩ := rfl @@ -445,7 +436,7 @@ theorem val_add_eq_ite {n : ℕ} (a b : Fin n) : section OfNatCoe @[simp] -theorem ofNat''_eq_cast (n : ℕ) [NeZero n] (a : ℕ) : (Fin.ofNat'' a : Fin n) = a := +theorem ofNat'_eq_cast (n : ℕ) [NeZero n] (a : ℕ) : Fin.ofNat' n a = a := rfl @[simp] lemma val_natCast (a n : ℕ) [NeZero n] : (a : Fin n).val = a % n := rfl From 9a958e834a5a2be5d704777f1e7dcf7c728f0888 Mon Sep 17 00:00:00 2001 From: Riccardo Brasca Date: Wed, 16 Oct 2024 06:24:48 +0000 Subject: [PATCH 004/505] doc: clarify doc of `Module.finrank` (#17780) --- Mathlib/LinearAlgebra/Dimension/Finrank.lean | 3 +++ 1 file changed, 3 insertions(+) diff --git a/Mathlib/LinearAlgebra/Dimension/Finrank.lean b/Mathlib/LinearAlgebra/Dimension/Finrank.lean index 69668ef8b667a..b702b49f38f5c 100644 --- a/Mathlib/LinearAlgebra/Dimension/Finrank.lean +++ b/Mathlib/LinearAlgebra/Dimension/Finrank.lean @@ -48,6 +48,9 @@ Defined by convention to be `0` if the space has infinite rank. For a vector space `M` over a field `R`, this is the same as the finite dimension of `M` over `R`. + +Note that it is possible to have `M` with `¬(Module.Finite R M)` but `finrank R M ≠ 0`, for example +`ℤ × ℚ/ℤ` has `finrank` equal to `1`. -/ noncomputable def finrank (R M : Type*) [Semiring R] [AddCommGroup M] [Module R M] : ℕ := Cardinal.toNat (Module.rank R M) From e398204ea3e5001b88c0c7f48e736664a9f0a235 Mon Sep 17 00:00:00 2001 From: Moritz Firsching Date: Wed, 16 Oct 2024 07:57:08 +0000 Subject: [PATCH 005/505] feat(NumberTheory/Fermat): fermat_eq_fermat_sub_one_pow_two_add_one (#17611) - [x] depends on: #17618 Co-authored-by: Moritz Firsching --- Mathlib/NumberTheory/Fermat.lean | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/Mathlib/NumberTheory/Fermat.lean b/Mathlib/NumberTheory/Fermat.lean index 0fcd105b0a600..5fe9b6b32bc73 100644 --- a/Mathlib/NumberTheory/Fermat.lean +++ b/Mathlib/NumberTheory/Fermat.lean @@ -58,6 +58,10 @@ theorem fermatNumber_eq_prod_add_two (n : ℕ) : rw [fermatNumber_product, Nat.sub_add_cancel] exact le_of_lt <| two_lt_fermatNumber _ +theorem fermatNumber_succ (n : ℕ) : fermatNumber (n + 1) = (fermatNumber n - 1) ^ 2 + 1 := by + rw [fermatNumber, pow_succ, mul_comm, Nat.pow_mul'] + rfl + /-- **Goldbach's theorem** : no two distinct Fermat numbers share a common factor greater than one. From 765ccc9135f3ebeec047798d6c715dc3224c0c73 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Wed, 16 Oct 2024 07:57:09 +0000 Subject: [PATCH 006/505] chore: update overview to use Std.HashMap, not Batteries.HashMap (#17797) --- docs/overview.yaml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/docs/overview.yaml b/docs/overview.yaml index d020bc5328e44..7113041d2dfc9 100644 --- a/docs/overview.yaml +++ b/docs/overview.yaml @@ -468,7 +468,7 @@ Data structures: Maps: key-value map: 'AList' red-black map: 'Batteries.RBMap' - hash map: 'Batteries.HashMap' + hash map: 'Std.HashMap' finitely supported function: 'Finsupp' finite map: 'Finmap' From f38babeaf882777867f57167bb35d3695d1c513f Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 16 Oct 2024 08:26:02 +0000 Subject: [PATCH 007/505] feat: `comap` of a finite measure along a measurable equiv is finite (#17639) From PFR --- Mathlib/MeasureTheory/Measure/MeasureSpace.lean | 4 ++++ Mathlib/MeasureTheory/Measure/Typeclasses.lean | 11 +++++++++++ 2 files changed, 15 insertions(+) diff --git a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean index 2bbeba4759747..6068dcd58e274 100644 --- a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean +++ b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean @@ -1029,6 +1029,10 @@ instance [NeZero μ] : NeZero (μ univ) := ⟨measure_univ_ne_zero.2 <| NeZero.n theorem measure_univ_pos : 0 < μ univ ↔ μ ≠ 0 := pos_iff_ne_zero.trans measure_univ_ne_zero +lemma nonempty_of_neZero (μ : Measure α) [NeZero μ] : Nonempty α := + (isEmpty_or_nonempty α).resolve_left fun h ↦ by + simpa [eq_empty_of_isEmpty] using NeZero.ne (μ univ) + /-! ### Pushforward and pullback -/ diff --git a/Mathlib/MeasureTheory/Measure/Typeclasses.lean b/Mathlib/MeasureTheory/Measure/Typeclasses.lean index fb5f43b48b9f2..e84781a13f059 100644 --- a/Mathlib/MeasureTheory/Measure/Typeclasses.lean +++ b/Mathlib/MeasureTheory/Measure/Typeclasses.lean @@ -122,6 +122,14 @@ theorem Measure.isFiniteMeasure_map {m : MeasurableSpace α} (μ : Measure α) [ · rw [map_of_not_aemeasurable hf] exact MeasureTheory.isFiniteMeasureZero +instance IsFiniteMeasure_comap (f : β → α) [IsFiniteMeasure μ] : IsFiniteMeasure (μ.comap f) where + measure_univ_lt_top := by + by_cases hf : Injective f ∧ ∀ s, MeasurableSet s → NullMeasurableSet (f '' s) μ + · rw [Measure.comap_apply₀ _ _ hf.1 hf.2 MeasurableSet.univ.nullMeasurableSet] + exact measure_lt_top μ _ + · rw [Measure.comap, dif_neg hf] + exact zero_lt_top + @[simp] theorem measureUnivNNReal_eq_zero [IsFiniteMeasure μ] : measureUnivNNReal μ = 0 ↔ μ = 0 := by rw [← MeasureTheory.Measure.measure_univ_eq_zero, ← coe_measureUnivNNReal] @@ -277,6 +285,9 @@ theorem isProbabilityMeasure_map {f : α → β} (hf : AEMeasurable f μ) : IsProbabilityMeasure (map f μ) := ⟨by simp [map_apply_of_aemeasurable, hf]⟩ +instance IsProbabilityMeasure_comap_equiv (f : β ≃ᵐ α) : IsProbabilityMeasure (μ.comap f) := by + rw [← MeasurableEquiv.map_symm]; exact isProbabilityMeasure_map f.symm.measurable.aemeasurable + /-- Note that this is not quite as useful as it looks because the measure takes values in `ℝ≥0∞`. Thus the subtraction appearing is the truncated subtraction of `ℝ≥0∞`, rather than the better-behaved subtraction of `ℝ`. -/ From b913345e5847bd22f23ab83b6256654095f161eb Mon Sep 17 00:00:00 2001 From: Robin Carlier <57142648+robin-carlier@users.noreply.github.com> Date: Wed, 16 Oct 2024 08:40:45 +0000 Subject: [PATCH 008/505] feat(CategoryTheory/Limits): definition of sifted categories (#17779) Introduce the class of sifted categories. A category `IsSifted` if it is nonempty and the diagonal functor is a final functor. - Show that sifted categories are stable under equivalences of categories. - Show that a category is sifted if and only if a small model of it is sifted. - Show that a sifted category is connected. - Show that a category with binary coproducts is sifted. --- Mathlib.lean | 1 + Mathlib/CategoryTheory/Limits/Sifted.lean | 104 ++++++++++++++++++++++ docs/references.bib | 11 +++ 3 files changed, 116 insertions(+) create mode 100644 Mathlib/CategoryTheory/Limits/Sifted.lean diff --git a/Mathlib.lean b/Mathlib.lean index e2a1e20cdd35a..4145efbc9c714 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1708,6 +1708,7 @@ import Mathlib.CategoryTheory.Limits.Shapes.WideEqualizers import Mathlib.CategoryTheory.Limits.Shapes.WidePullbacks import Mathlib.CategoryTheory.Limits.Shapes.ZeroMorphisms import Mathlib.CategoryTheory.Limits.Shapes.ZeroObjects +import Mathlib.CategoryTheory.Limits.Sifted import Mathlib.CategoryTheory.Limits.SmallComplete import Mathlib.CategoryTheory.Limits.Types import Mathlib.CategoryTheory.Limits.TypesFiltered diff --git a/Mathlib/CategoryTheory/Limits/Sifted.lean b/Mathlib/CategoryTheory/Limits/Sifted.lean new file mode 100644 index 0000000000000..0949b037550c1 --- /dev/null +++ b/Mathlib/CategoryTheory/Limits/Sifted.lean @@ -0,0 +1,104 @@ +/- +Copyright (c) 2024 Robin Carlier. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Robin Carlier +-/ +import Mathlib.CategoryTheory.Limits.Final +/-! +# Sifted categories + +A category `C` is sifted if `C` is nonempty and the diagonal functor `C ⥤ C × C` is final. +Sifted categories can be characterized as those such that the colimit functor `(C ⥤ Type) ⥤ Type ` +preserves finite products. + +## Main results +- `isSifted_of_hasBinaryCoproducts_and_nonempty`: A nonempty category with binary coproducts is + sifted. + +## References +- [nLab, *Sifted category*](https://ncatlab.org/nlab/show/sifted+category) +- [*Algebraic Theories*, Chapter 2.][Adámek_Rosický_Vitale_2010] +-/ + +universe w v v₁ u u₁ + +namespace CategoryTheory + +open Limits Functor + +section + +variable (C : Type u) [Category.{v} C] + +/-- A category `C` `IsSiftedOrEmpty` if the diagonal functor `C ⥤ C × C` is final. -/ +abbrev IsSiftedOrEmpty : Prop := Final (diag C) + +/-- A category `C` `IsSfited` if +1. the diagonal functor `C ⥤ C × C` is final. +2. there exists some object. -/ +class IsSifted extends IsSiftedOrEmpty C : Prop where + [nonempty : Nonempty C] + +attribute [instance] IsSifted.nonempty + +namespace IsSifted + +variable {C} + +/-- Being sifted is preserved by equivalences of categories -/ +lemma isSifted_of_equiv [IsSifted C] {D : Type u₁} [Category.{v₁} D] (e : D ≌ C) : IsSifted D := + letI : Final (diag D) := by + letI : D × D ≌ C × C:= Equivalence.prod e e + have sq : (e.inverse ⋙ diag D ⋙ this.functor ≅ diag C) := + NatIso.ofComponents (fun c ↦ by dsimp [this] + exact Iso.prod (e.counitIso.app c) (e.counitIso.app c)) + apply_rules [final_iff_comp_equivalence _ this.functor|>.mpr, + final_iff_final_comp e.inverse _ |>.mpr, final_of_natIso sq.symm] + letI : _root_.Nonempty D := ⟨e.inverse.obj (_root_.Nonempty.some IsSifted.nonempty)⟩ + ⟨⟩ + +/-- In particular a category is sifted iff and only if it is so when viewed as a small category -/ +lemma isSifted_iff_asSmallIsSifted : IsSifted C ↔ IsSifted (AsSmall.{w} C) where + mp _ := isSifted_of_equiv AsSmall.equiv.symm + mpr _ := isSifted_of_equiv AsSmall.equiv + +/-- A sifted category is connected. -/ +instance [IsSifted C] : IsConnected C := + isConnected_of_zigzag + (by intro c₁ c₂ + have X : StructuredArrow (c₁, c₂) (diag C) := + letI S : Final (diag C) := by infer_instance + Nonempty.some (S.out (c₁, c₂)).is_nonempty + use [X.right, c₂] + constructor + · constructor + · exact Zag.of_hom X.hom.fst + · simp + exact Zag.of_inv X.hom.snd + · rfl) + +/-- A category with binary coproducts is sifted or empty. -/ +instance [HasBinaryCoproducts C] : IsSiftedOrEmpty C := by + constructor + rintro ⟨c₁, c₂⟩ + haveI : _root_.Nonempty <| StructuredArrow (c₁,c₂) (diag C) := + ⟨.mk ((coprod.inl : c₁ ⟶ c₁ ⨿ c₂), (coprod.inr : c₂ ⟶ c₁ ⨿ c₂))⟩ + apply isConnected_of_zigzag + rintro ⟨_, c, f⟩ ⟨_, c', g⟩ + dsimp only [const_obj_obj, diag_obj, prod_Hom] at f g + use [.mk ((coprod.inl : c₁ ⟶ c₁ ⨿ c₂), (coprod.inr : c₂ ⟶ c₁ ⨿ c₂)), .mk (g.fst, g.snd)] + simp only [colimit.cocone_x, diag_obj, Prod.mk.eta, List.chain_cons, List.Chain.nil, and_true, + ne_eq, reduceCtorEq, not_false_eq_true, List.getLast_cons, List.cons_ne_self, + List.getLast_singleton] + exact ⟨⟨Zag.of_inv <| StructuredArrow.homMk <| coprod.desc f.fst f.snd, + Zag.of_hom <| StructuredArrow.homMk <| coprod.desc g.fst g.snd⟩, rfl⟩ + +/-- A nonempty category with binary coproducts is sifted. -/ +instance isSifted_of_hasBinaryCoproducts_and_nonempty [_root_.Nonempty C] [HasBinaryCoproducts C] : + IsSifted C where + +end IsSifted + +end + +end CategoryTheory diff --git a/docs/references.bib b/docs/references.bib index 67919e2ced15f..39e500681a649 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -34,6 +34,17 @@ @InProceedings{ adhesive2004 url = {https://www.ioc.ee/~pawel/papers/adhesive.pdf} } +@Book{ Adámek_Rosický_Vitale_2010, + place = {Cambridge}, + series = {Cambridge Tracts in Mathematics}, + title = {Algebraic Theories: A Categorical Introduction to General + Algebra}, + publisher = {Cambridge University Press}, + author = {Adámek, J. and Rosický, J. and Vitale, E. M.}, + year = {2010}, + collection = {Cambridge Tracts in Mathematics} +} + @Article{ ahrens2017, author = {Benedikt Ahrens and Peter LeFanu Lumsdaine}, year = {2019}, From 4564def38d31320ae5385d4d4463f4942bcbafc3 Mon Sep 17 00:00:00 2001 From: Quang Dao Date: Wed, 16 Oct 2024 08:50:02 +0000 Subject: [PATCH 009/505] chore(Multiset/Basic): added `gcongr` and `congr` to `Multiset` (#17793) While working on some theorems about `MvPolynomial.degrees`, I noticed that a lot of `Multiset` functions do not have `gcongr` or `congr` attributes. So I added them in this PR. Note: I only looked at the definitions in `Multiset/Basic`. If this PR is approved I will do the same for other functions in `Multiset` folder. Co-authored-by: Quang Dao --- Mathlib/Data/Multiset/Basic.lean | 19 ++++++++++++++++--- 1 file changed, 16 insertions(+), 3 deletions(-) diff --git a/Mathlib/Data/Multiset/Basic.lean b/Mathlib/Data/Multiset/Basic.lean index b54ec18a4e531..7dc4316f0e993 100644 --- a/Mathlib/Data/Multiset/Basic.lean +++ b/Mathlib/Data/Multiset/Basic.lean @@ -864,6 +864,10 @@ theorem nsmul_singleton (a : α) (n) : n • ({a} : Multiset α) = replicate n a theorem replicate_le_replicate (a : α) {k n : ℕ} : replicate k a ≤ replicate n a ↔ k ≤ n := _root_.trans (by rw [← replicate_le_coe, coe_replicate]) (List.replicate_sublist_replicate a) +@[gcongr] +theorem replicate_mono (a : α) {k n : ℕ} (h : k ≤ n) : replicate k a ≤ replicate n a := + (replicate_le_replicate a).2 h + theorem le_replicate_iff {m : Multiset α} {a : α} {n : ℕ} : m ≤ replicate n a ↔ ∃ k ≤ n, m = replicate k a := ⟨fun h => ⟨card m, (card_mono h).trans_eq (card_replicate _ _), @@ -972,6 +976,7 @@ theorem erase_comm (s : Multiset α) (a b : α) : (s.erase a).erase b = (s.erase instance : RightCommutative erase (α := α) := ⟨erase_comm⟩ +@[gcongr] theorem erase_le_erase {s t : Multiset α} (a : α) (h : s ≤ t) : s.erase a ≤ t.erase a := leInductionOn h fun h => (h.erase _).subperm @@ -1151,11 +1156,11 @@ theorem eq_of_mem_map_const {b₁ b₂ : β} {l : List α} (h : b₁ ∈ map (Fu b₁ = b₂ := eq_of_mem_replicate (n := card (l : Multiset α)) <| by rwa [map_const] at h -@[simp] +@[simp, gcongr] theorem map_le_map {f : α → β} {s t : Multiset α} (h : s ≤ t) : map f s ≤ map f t := leInductionOn h fun h => (h.map f).subperm -@[simp] +@[simp, gcongr] theorem map_lt_map {f : α → β} {s t : Multiset α} (h : s < t) : s.map f < t.map f := by refine (map_le_map h.le).lt_of_not_le fun H => h.ne <| eq_of_le_of_card_le h.le ?_ rw [← s.card_map f, ← t.card_map f] @@ -1165,7 +1170,7 @@ theorem map_mono (f : α → β) : Monotone (map f) := fun _ _ => map_le_map theorem map_strictMono (f : α → β) : StrictMono (map f) := fun _ _ => map_lt_map -@[simp] +@[simp, gcongr] theorem map_subset_map {f : α → β} {s t : Multiset α} (H : s ⊆ t) : map f s ⊆ map f t := fun _b m => let ⟨a, h, e⟩ := mem_map.1 m mem_map.2 ⟨a, H h, e⟩ @@ -1521,6 +1526,7 @@ theorem le_union_right (s t : Multiset α) : t ≤ s ∪ t := theorem eq_union_left : t ≤ s → s ∪ t = s := tsub_add_cancel_of_le +@[gcongr] theorem union_le_union_right (h : s ≤ t) (u) : s ∪ u ≤ t ∪ u := add_le_add_right (tsub_le_tsub_right h _) u @@ -1626,6 +1632,7 @@ theorem inter_comm (s t : Multiset α) : s ∩ t = t ∩ s := inf_comm _ _ theorem eq_union_right (h : s ≤ t) : s ∪ t = t := by rw [union_comm, eq_union_left h] +@[gcongr] theorem union_le_union_left (h : s ≤ t) (u) : u ∪ s ≤ u ∪ t := sup_le_sup_left h _ @@ -1710,6 +1717,7 @@ theorem filter_zero : filter p 0 = 0 := Please re-enable the linter once we moved to `nightly-2024-06-22` or later. -/ set_option linter.deprecated false in +@[congr] theorem filter_congr {p q : α → Prop} [DecidablePred p] [DecidablePred q] {s : Multiset α} : (∀ x ∈ s, p x ↔ q x) → filter p s = filter q s := Quot.inductionOn s fun _l h => congr_arg ofList <| filter_congr' <| by simpa using h @@ -1726,6 +1734,7 @@ theorem filter_le (s : Multiset α) : filter p s ≤ s := theorem filter_subset (s : Multiset α) : filter p s ⊆ s := subset_of_le <| filter_le _ _ +@[gcongr] theorem filter_le_filter {s t} (h : s ≤ t) : filter p s ≤ filter p t := leInductionOn h fun h => (h.filter (p ·)).subperm @@ -1938,6 +1947,7 @@ theorem map_filterMap_of_inv (f : α → Option β) (g : β → α) (H : ∀ x : (s : Multiset α) : map g (filterMap f s) = s := Quot.inductionOn s fun l => congr_arg ofList <| List.map_filterMap_of_inv f g H l +@[gcongr] theorem filterMap_le_filterMap (f : α → Option β) {s t : Multiset α} (h : s ≤ t) : filterMap f s ≤ filterMap f t := leInductionOn h fun h => (h.filterMap _).subperm @@ -2006,6 +2016,7 @@ theorem countP_sub [DecidableEq α] {s t : Multiset α} (h : t ≤ s) : countP p (s - t) = countP p s - countP p t := by simp [countP_eq_card_filter, h, filter_le_filter] +@[gcongr] theorem countP_le_of_le {s t} (h : s ≤ t) : countP p s ≤ countP p t := by simpa [countP_eq_card_filter] using card_le_card (filter_le_filter p h) @@ -2065,6 +2076,7 @@ theorem countP_eq_card {s} : countP p s = card s ↔ ∀ a ∈ s, p a := theorem countP_pos_of_mem {s a} (h : a ∈ s) (pa : p a) : 0 < countP p s := countP_pos.2 ⟨_, h, pa⟩ +@[congr] theorem countP_congr {s s' : Multiset α} (hs : s = s') {p p' : α → Prop} [DecidablePred p] [DecidablePred p'] (hp : ∀ x ∈ s, p x = p' x) : s.countP p = s'.countP p' := by @@ -2108,6 +2120,7 @@ theorem count_cons_of_ne {a b : α} (h : a ≠ b) (s : Multiset α) : count a (b theorem count_le_card (a : α) (s) : count a s ≤ card s := countP_le_card _ _ +@[gcongr] theorem count_le_of_le (a : α) {s t} : s ≤ t → count a s ≤ count a t := countP_le_of_le _ From 8e7833a57c7331fcf2a2b9776482ae29dca6d004 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 16 Oct 2024 10:11:37 +0000 Subject: [PATCH 010/505] feat: `MeasurableSpace` instance for the quotient of a module by a submodule (#17812) From PFR --- Mathlib.lean | 1 + .../Constructions/SubmoduleQuotient.lean | 20 +++++++++++++++++++ .../MeasureTheory/MeasurableSpace/Basic.lean | 9 +++++++++ 3 files changed, 30 insertions(+) create mode 100644 Mathlib/MeasureTheory/Constructions/SubmoduleQuotient.lean diff --git a/Mathlib.lean b/Mathlib.lean index 4145efbc9c714..ad61f598310a6 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3241,6 +3241,7 @@ import Mathlib.MeasureTheory.Constructions.Polish.EmbeddingReal import Mathlib.MeasureTheory.Constructions.Prod.Basic import Mathlib.MeasureTheory.Constructions.Prod.Integral import Mathlib.MeasureTheory.Constructions.Projective +import Mathlib.MeasureTheory.Constructions.SubmoduleQuotient import Mathlib.MeasureTheory.Constructions.UnitInterval import Mathlib.MeasureTheory.Covering.Besicovitch import Mathlib.MeasureTheory.Covering.BesicovitchVectorSpace diff --git a/Mathlib/MeasureTheory/Constructions/SubmoduleQuotient.lean b/Mathlib/MeasureTheory/Constructions/SubmoduleQuotient.lean new file mode 100644 index 0000000000000..e5bcc2ccd92dd --- /dev/null +++ b/Mathlib/MeasureTheory/Constructions/SubmoduleQuotient.lean @@ -0,0 +1,20 @@ +/- +Copyright (c) 2024 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies +-/ +import Mathlib.LinearAlgebra.Quotient +import Mathlib.MeasureTheory.MeasurableSpace.Basic + +/-! +# Measurability on the quotient of a module by a submodule +-/ + +namespace Submodule.Quotient +variable {R M : Type*} [Ring R] [AddCommGroup M] [Module R M] {p : Submodule R M} + +instance [MeasurableSpace M] : MeasurableSpace (M ⧸ p) := Quotient.instMeasurableSpace +instance [MeasurableSpace M] [DiscreteMeasurableSpace M] : DiscreteMeasurableSpace (M ⧸ p) := + Quotient.instDiscreteMeasurableSpace + +end Submodule.Quotient diff --git a/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean b/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean index 573857eeaad27..49c7a06029ead 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean @@ -472,6 +472,15 @@ nonrec theorem QuotientGroup.measurable_from_quotient {G} [Group G] [MeasurableS {S : Subgroup G} {f : G ⧸ S → α} : Measurable f ↔ Measurable (f ∘ ((↑) : G → G ⧸ S)) := measurable_from_quotient +instance Quotient.instDiscreteMeasurableSpace {α} {s : Setoid α} [MeasurableSpace α] + [DiscreteMeasurableSpace α] : DiscreteMeasurableSpace (Quotient s) where + forall_measurableSet _ := measurableSet_quotient.2 .of_discrete + +@[to_additive] +instance QuotientGroup.instDiscreteMeasurableSpace {G} [Group G] [MeasurableSpace G] + [DiscreteMeasurableSpace G] (S : Subgroup G) : DiscreteMeasurableSpace (G ⧸ S) := + Quotient.instDiscreteMeasurableSpace + end Quotient section Subtype From 88e7113ac302a91f873cb154ca186f8a0c7b89ff Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Wed, 16 Oct 2024 10:27:00 +0000 Subject: [PATCH 011/505] =?UTF-8?q?feat(SetTheory/Ordinal/Arithmetic):=20`?= =?UTF-8?q?a=20=E2=89=A4=20b=20=E2=86=92=20a=20/=20c=20=E2=89=A4=20b=20/?= =?UTF-8?q?=20c`=20(#17741)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- Mathlib/SetTheory/Ordinal/Arithmetic.lean | 6 ++++++ 1 file changed, 6 insertions(+) diff --git a/Mathlib/SetTheory/Ordinal/Arithmetic.lean b/Mathlib/SetTheory/Ordinal/Arithmetic.lean index 65ed7c402f75f..aa95e0d9ce216 100644 --- a/Mathlib/SetTheory/Ordinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Ordinal/Arithmetic.lean @@ -865,6 +865,12 @@ theorem zero_div (a : Ordinal) : 0 / a = 0 := theorem mul_div_le (a b : Ordinal) : b * (a / b) ≤ a := if b0 : b = 0 then by simp only [b0, zero_mul, Ordinal.zero_le] else (le_div b0).1 le_rfl +theorem div_le_left {a b : Ordinal} (h : a ≤ b) (c : Ordinal) : a / c ≤ b / c := by + obtain rfl | hc := eq_or_ne c 0 + · rw [div_zero, div_zero] + · rw [le_div hc] + exact (mul_div_le a c).trans h + theorem mul_add_div (a) {b : Ordinal} (b0 : b ≠ 0) (c) : (b * a + c) / b = a + c / b := by apply le_antisymm · apply (div_le b0).2 From 8591fb6789e584a8f8d46509cc803b7475140a93 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 16 Oct 2024 10:27:01 +0000 Subject: [PATCH 012/505] =?UTF-8?q?feat:=20`s=20=E2=8A=86=20f=20''=20t=20?= =?UTF-8?q?=E2=86=92=20f=20=E2=81=BB=C2=B9'=20s=20=E2=8A=86=20t`=20(#17788?= =?UTF-8?q?)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- Mathlib/Data/Finset/Preimage.lean | 4 ++++ Mathlib/Data/Set/Image.lean | 6 ++++++ 2 files changed, 10 insertions(+) diff --git a/Mathlib/Data/Finset/Preimage.lean b/Mathlib/Data/Finset/Preimage.lean index 5a78ba233b2d2..1331695180e29 100644 --- a/Mathlib/Data/Finset/Preimage.lean +++ b/Mathlib/Data/Finset/Preimage.lean @@ -95,6 +95,10 @@ theorem image_preimage_of_bij [DecidableEq β] (f : α → β) (s : Finset β) (hf : Set.BijOn f (f ⁻¹' ↑s) ↑s) : image f (preimage s f hf.injOn) = s := Finset.coe_inj.1 <| by simpa using hf.image_eq +lemma preimage_subset_of_subset_image [DecidableEq β] {f : α → β} {s : Finset β} {t : Finset α} + (hs : s ⊆ t.image f) {hf} : s.preimage f hf ⊆ t := by + rw [← coe_subset, coe_preimage]; exact Set.preimage_subset (mod_cast hs) hf + theorem preimage_subset {f : α ↪ β} {s : Finset β} {t : Finset α} (hs : s ⊆ t.map f) : s.preimage f f.injective.injOn ⊆ t := fun _ h => (mem_map' f).1 (hs (mem_preimage.1 h)) diff --git a/Mathlib/Data/Set/Image.lean b/Mathlib/Data/Set/Image.lean index e70bfdc19f8d5..bbd73befb9a8c 100644 --- a/Mathlib/Data/Set/Image.lean +++ b/Mathlib/Data/Set/Image.lean @@ -168,6 +168,12 @@ theorem preimage_subtype_coe_eq_compl {s u v : Set α} (hsuv : s ⊆ u ∪ v) · intro hx exact Or.elim (hsuv x_in_s) id fun hx' => hx.elim hx' +lemma preimage_subset {s t} (hs : s ⊆ f '' t) (hf : Set.InjOn f (f ⁻¹' s)) : f ⁻¹' s ⊆ t := by + rintro a ha + obtain ⟨b, hb, hba⟩ := hs ha + rwa [hf ha _ hba.symm] + simpa [hba] + end Preimage /-! ### Image of a set under a function -/ From 4fde5a554745253ba37fb31dbd63eda984eadc58 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Wed, 16 Oct 2024 10:27:02 +0000 Subject: [PATCH 013/505] feat(SetTheory/Ordinal/Exponential): more lemmas on `opow` (#17804) --- Mathlib/SetTheory/Ordinal/Exponential.lean | 9 ++++++++- 1 file changed, 8 insertions(+), 1 deletion(-) diff --git a/Mathlib/SetTheory/Ordinal/Exponential.lean b/Mathlib/SetTheory/Ordinal/Exponential.lean index 902f6a296c035..3fad6ff39bde5 100644 --- a/Mathlib/SetTheory/Ordinal/Exponential.lean +++ b/Mathlib/SetTheory/Ordinal/Exponential.lean @@ -157,6 +157,9 @@ theorem opow_le_opow_left {a b : Ordinal} (c : Ordinal) (ab : a ≤ b) : a ^ c (opow_le_of_limit a0 l).2 fun b' h => (IH _ h).trans (opow_le_opow_right ((Ordinal.pos_iff_ne_zero.2 a0).trans_le ab) h.le) +theorem opow_le_opow {a b c d : Ordinal} (hac : a ≤ c) (hbd : b ≤ d) (hc : 0 < c) : a ^ b ≤ c ^ d := + (opow_le_opow_left b hac).trans (opow_le_opow_right hc hbd) + theorem left_le_opow (a : Ordinal) {b : Ordinal} (b1 : 0 < b) : a ≤ a ^ b := by nth_rw 1 [← opow_one a] cases' le_or_gt a 1 with a1 a1 @@ -167,8 +170,12 @@ theorem left_le_opow (a : Ordinal) {b : Ordinal} (b1 : 0 < b) : a ≤ a ^ b := b rw [a1, one_opow, one_opow] rwa [opow_le_opow_iff_right a1, one_le_iff_pos] +theorem left_lt_opow {a b : Ordinal} (ha : 1 < a) (hb : 1 < b) : a < a ^ b := by + conv_lhs => rw [← opow_one a] + rwa [opow_lt_opow_iff_right ha] + theorem right_le_opow {a : Ordinal} (b : Ordinal) (a1 : 1 < a) : b ≤ a ^ b := - (isNormal_opow a1).id_le _ + (isNormal_opow a1).le_apply theorem opow_lt_opow_left_of_succ {a b c : Ordinal} (ab : a < b) : a ^ succ c < b ^ succ c := by rw [opow_succ, opow_succ] From 36c03bdc010a49587fe6b786f92f6eb1e9a75040 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Wed, 16 Oct 2024 11:05:15 +0000 Subject: [PATCH 014/505] refactor(AlgebraicGeometry): Make `Scheme.Hom` a structure. (#15093) We make `LocallyRingedSpace.Hom` (resp. `Scheme.Hom`) into a structure instead of a subtype (resp. type-synonym of the former) to make clearer the abstraction boundaries between them and presheafed spaces, which allow us to remove some `erw`s. Plus, this enables us to write `f.base` and `f.c` instead of `f.val.base` and `f.val.c`. Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- Mathlib/AlgebraicGeometry/AffineScheme.lean | 54 ++--- Mathlib/AlgebraicGeometry/Cover/Open.lean | 48 ++--- Mathlib/AlgebraicGeometry/FunctionField.lean | 4 +- .../GammaSpecAdjunction.lean | 60 +++--- Mathlib/AlgebraicGeometry/Gluing.lean | 42 ++-- Mathlib/AlgebraicGeometry/Limits.lean | 35 ++-- .../AlgebraicGeometry/Morphisms/Affine.lean | 4 +- .../AlgebraicGeometry/Morphisms/Basic.lean | 2 +- .../Morphisms/ClosedImmersion.lean | 16 +- .../Morphisms/Constructors.lean | 20 +- .../AlgebraicGeometry/Morphisms/IsIso.lean | 4 +- .../Morphisms/OpenImmersion.lean | 4 +- .../Morphisms/Preimmersion.lean | 6 +- .../Morphisms/QuasiCompact.lean | 18 +- .../Morphisms/QuasiSeparated.lean | 2 +- .../Morphisms/RingHomProperties.lean | 2 +- .../Morphisms/UnderlyingMap.lean | 6 +- .../Morphisms/UniversallyClosed.lean | 4 +- Mathlib/AlgebraicGeometry/Noetherian.lean | 2 +- Mathlib/AlgebraicGeometry/OpenImmersion.lean | 139 ++++++------- .../ProjectiveSpectrum/Scheme.lean | 34 +-- Mathlib/AlgebraicGeometry/Properties.lean | 14 +- Mathlib/AlgebraicGeometry/Pullbacks.lean | 20 +- Mathlib/AlgebraicGeometry/ResidueField.lean | 22 +- Mathlib/AlgebraicGeometry/Restrict.lean | 36 ++-- Mathlib/AlgebraicGeometry/Scheme.lean | 195 ++++++++++-------- Mathlib/AlgebraicGeometry/Spec.lean | 10 +- .../RingedSpace/LocallyRingedSpace.lean | 152 ++++++++------ .../LocallyRingedSpace/HasColimits.lean | 109 +++++----- .../LocallyRingedSpace/ResidueField.lean | 32 +-- .../Geometry/RingedSpace/OpenImmersion.lean | 68 +++--- .../Geometry/RingedSpace/PresheafedSpace.lean | 3 + .../RingedSpace/PresheafedSpace/Gluing.lean | 2 +- 33 files changed, 608 insertions(+), 561 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/AffineScheme.lean b/Mathlib/AlgebraicGeometry/AffineScheme.lean index 35d58412138ae..731a32691323c 100644 --- a/Mathlib/AlgebraicGeometry/AffineScheme.lean +++ b/Mathlib/AlgebraicGeometry/AffineScheme.lean @@ -281,12 +281,12 @@ def isoSpec : (X.presheaf.mapIso (eqToIso U.openEmbedding_obj_top).op).op open LocalRing in -lemma isoSpec_hom_val_base_apply (x : U) : - hU.isoSpec.hom.1.base x = (Spec.map (X.presheaf.germ _ x x.2)).val.base (closedPoint _) := by - dsimp [IsAffineOpen.isoSpec_hom, Scheme.isoSpec_hom, Scheme.toSpecΓ_val_base] - rw [← Scheme.comp_val_base_apply, ← Spec.map_comp, +lemma isoSpec_hom_base_apply (x : U) : + hU.isoSpec.hom.base x = (Spec.map (X.presheaf.germ _ x x.2)).base (closedPoint _) := by + dsimp [IsAffineOpen.isoSpec_hom, Scheme.isoSpec_hom, Scheme.toSpecΓ_base] + rw [← Scheme.comp_base_apply, ← Spec.map_comp, (Iso.eq_comp_inv _).mpr (Scheme.Opens.germ_stalkIso_hom U (V := ⊤) x trivial), - X.presheaf.germ_res_assoc, Spec.map_comp, Scheme.comp_val_base_apply] + X.presheaf.germ_res_assoc, Spec.map_comp, Scheme.comp_base_apply] congr 1 have := isLocalRingHom_of_isIso (U.stalkIso x).inv exact LocalRing.comap_closedPoint (U.stalkIso x).inv @@ -322,7 +322,7 @@ lemma isoSpec_inv_ι : hU.isoSpec.inv ≫ U.ι = hU.fromSpec := rfl @[simp] theorem range_fromSpec : - Set.range hU.fromSpec.1.base = (U : Set X) := by + Set.range hU.fromSpec.base = (U : Set X) := by delta IsAffineOpen.fromSpec; dsimp [IsAffineOpen.isoSpec_inv] rw [Set.range_comp, Set.range_iff_surjective.mpr, Set.image_univ] · exact Subtype.range_coe @@ -374,7 +374,7 @@ lemma fromSpec_app_of_le (V : X.Opens) (h : U ≤ V) : include hU in protected theorem isCompact : IsCompact (U : Set X) := by - convert @IsCompact.image _ _ _ _ Set.univ hU.fromSpec.1.base PrimeSpectrum.compactSpace.1 + convert @IsCompact.image _ _ _ _ Set.univ hU.fromSpec.base PrimeSpectrum.compactSpace.1 (by fun_prop) convert hU.range_fromSpec.symm exact Set.image_univ @@ -398,7 +398,7 @@ theorem _root_.AlgebraicGeometry.Scheme.Hom.isAffineOpen_iff_of_isOpenImmersion refine ⟨fun hU => @isAffine_of_isIso _ _ (IsOpenImmersion.isoOfRangeEq (X.ofRestrict U.openEmbedding ≫ f) (Y.ofRestrict _) ?_).hom ?_ hU, fun hU => hU.image_of_isOpenImmersion f⟩ - · rw [Scheme.comp_val_base, coe_comp, Set.range_comp] + · rw [Scheme.comp_base, coe_comp, Set.range_comp] dsimp [Opens.coe_inclusion', Scheme.restrict] erw [Subtype.range_coe, Subtype.range_coe] -- now `erw` after #13170 rfl @@ -467,7 +467,7 @@ theorem fromSpec_image_basicOpen : hU.fromSpec ''ᵁ (PrimeSpectrum.basicOpen f) = X.basicOpen f := by rw [← hU.fromSpec_preimage_basicOpen] ext1 - change hU.fromSpec.val.base '' (hU.fromSpec.val.base ⁻¹' (X.basicOpen f : Set X)) = _ + change hU.fromSpec.base '' (hU.fromSpec.base ⁻¹' (X.basicOpen f : Set X)) = _ rw [Set.image_preimage_eq_inter_range, Set.inter_eq_left, hU.range_fromSpec] exact Scheme.basicOpen_le _ _ @@ -519,7 +519,7 @@ this is the canonical map `Γ(𝒪ₓ, D(f)) ⟶ Γ(Spec 𝒪ₓ(U), D(f))` This is an isomorphism, as witnessed by an `IsIso` instance. -/ def basicOpenSectionsToAffine : Γ(X, X.basicOpen f) ⟶ Γ(Spec Γ(X, U), PrimeSpectrum.basicOpen f) := - hU.fromSpec.1.c.app (op <| X.basicOpen f) ≫ + hU.fromSpec.c.app (op <| X.basicOpen f) ≫ (Spec Γ(X, U)).presheaf.map (eqToHom <| (hU.fromSpec_preimage_basicOpen f).symm).op instance basicOpenSectionsToAffine_isIso : @@ -607,28 +607,28 @@ theorem _root_.AlgebraicGeometry.exists_basicOpen_le_affine_inter /-- The prime ideal of `𝒪ₓ(U)` corresponding to a point `x : U`. -/ noncomputable def primeIdealOf (x : U) : PrimeSpectrum Γ(X, U) := - hU.isoSpec.hom.1.base x + hU.isoSpec.hom.base x theorem fromSpec_primeIdealOf (x : U) : - hU.fromSpec.val.base (hU.primeIdealOf x) = x.1 := by + hU.fromSpec.base (hU.primeIdealOf x) = x.1 := by dsimp only [IsAffineOpen.fromSpec, Subtype.coe_mk, IsAffineOpen.primeIdealOf] - rw [← Scheme.comp_val_base_apply, Iso.hom_inv_id_assoc] + rw [← Scheme.comp_base_apply, Iso.hom_inv_id_assoc] rfl open LocalRing in theorem primeIdealOf_eq_map_closedPoint (x : U) : - hU.primeIdealOf x = (Spec.map (X.presheaf.germ _ x x.2)).val.base (closedPoint _) := - hU.isoSpec_hom_val_base_apply _ + hU.primeIdealOf x = (Spec.map (X.presheaf.germ _ x x.2)).base (closedPoint _) := + hU.isoSpec_hom_base_apply _ -theorem isLocalization_stalk' (y : PrimeSpectrum Γ(X, U)) (hy : hU.fromSpec.1.base y ∈ U) : +theorem isLocalization_stalk' (y : PrimeSpectrum Γ(X, U)) (hy : hU.fromSpec.base y ∈ U) : @IsLocalization.AtPrime (R := Γ(X, U)) - (S := X.presheaf.stalk <| hU.fromSpec.1.base y) _ _ + (S := X.presheaf.stalk <| hU.fromSpec.base y) _ _ ((TopCat.Presheaf.algebra_section_stalk X.presheaf _)) y.asIdeal _ := by apply (@IsLocalization.isLocalization_iff_of_ringEquiv (R := Γ(X, U)) - (S := X.presheaf.stalk (hU.fromSpec.1.base y)) _ y.asIdeal.primeCompl _ - (TopCat.Presheaf.algebra_section_stalk X.presheaf ⟨hU.fromSpec.1.base y, hy⟩) _ _ + (S := X.presheaf.stalk (hU.fromSpec.base y)) _ y.asIdeal.primeCompl _ + (TopCat.Presheaf.algebra_section_stalk X.presheaf ⟨hU.fromSpec.base y, hy⟩) _ _ (asIso <| hU.fromSpec.stalkMap y).commRingCatIsoToRingEquiv).mpr -- Porting note: need to know what the ring is and after convert, instead of equality -- we get an `iff`. @@ -646,7 +646,7 @@ theorem isLocalization_stalk (x : U) : IsLocalization.AtPrime (X.presheaf.stalk x) (hU.primeIdealOf x).asIdeal := by rcases x with ⟨x, hx⟩ set y := hU.primeIdealOf ⟨x, hx⟩ with hy - have : hU.fromSpec.val.base y = x := hy ▸ hU.fromSpec_primeIdealOf ⟨x, hx⟩ + have : hU.fromSpec.base y = x := hy ▸ hU.fromSpec_primeIdealOf ⟨x, hx⟩ clear_value y subst this exact hU.isLocalization_stalk' y hx @@ -666,11 +666,11 @@ theorem basicOpen_union_eq_self_iff (s : Set Γ(X, U)) : ⨆ f : s, X.basicOpen (f : Γ(X, U)) = U ↔ Ideal.span s = ⊤ := by trans ⋃ i : s, (PrimeSpectrum.basicOpen i.1).1 = Set.univ · trans - hU.fromSpec.1.base ⁻¹' (⨆ f : s, X.basicOpen (f : Γ(X, U))).1 = - hU.fromSpec.1.base ⁻¹' U.1 + hU.fromSpec.base ⁻¹' (⨆ f : s, X.basicOpen (f : Γ(X, U))).1 = + hU.fromSpec.base ⁻¹' U.1 · refine ⟨fun h => by rw [h], ?_⟩ intro h - apply_fun Set.image hU.fromSpec.1.base at h + apply_fun Set.image hU.fromSpec.base at h rw [Set.image_preimage_eq_inter_range, Set.image_preimage_eq_inter_range, hU.range_fromSpec] at h simp only [Set.inter_self, Opens.carrier_eq_coe, Set.inter_eq_right] at h @@ -765,7 +765,7 @@ section ZeroLocus /-- On a locally ringed space `X`, the preimage of the zero locus of the prime spectrum of `Γ(X, ⊤)` under `toΓSpecFun` agrees with the associated zero locus on `X`. -/ lemma Scheme.toΓSpec_preimage_zeroLocus_eq {X : Scheme.{u}} (s : Set Γ(X, ⊤)) : - X.toSpecΓ.val.base ⁻¹' PrimeSpectrum.zeroLocus s = X.zeroLocus s := + X.toSpecΓ.base ⁻¹' PrimeSpectrum.zeroLocus s = X.zeroLocus s := LocallyRingedSpace.toΓSpec_preimage_zeroLocus_eq s open ConcreteCategory @@ -774,9 +774,9 @@ open ConcreteCategory is the zero locus in terms of the prime spectrum of `Γ(X, ⊤)`. -/ lemma Scheme.toΓSpec_image_zeroLocus_eq_of_isAffine {X : Scheme.{u}} [IsAffine X] (s : Set Γ(X, ⊤)) : - X.isoSpec.hom.val.base '' X.zeroLocus s = PrimeSpectrum.zeroLocus s := by + X.isoSpec.hom.base '' X.zeroLocus s = PrimeSpectrum.zeroLocus s := by erw [← X.toΓSpec_preimage_zeroLocus_eq, Set.image_preimage_eq] - exact (bijective_of_isIso X.isoSpec.hom.val.base).surjective + exact (bijective_of_isIso X.isoSpec.hom.base).surjective /-- If `X` is an affine scheme, every closed set of `X` is the zero locus of a set of global sections. -/ @@ -788,7 +788,7 @@ lemma Scheme.eq_zeroLocus_of_isClosed_of_isAffine (X : Scheme.{u}) [IsAffine X] obtain ⟨I, (hI : Z = _)⟩ := (PrimeSpectrum.isClosed_iff_zeroLocus_ideal _).mp hZ use I simp only [← Scheme.toΓSpec_preimage_zeroLocus_eq, ← hI, Z] - erw [Set.preimage_image_eq _ (bijective_of_isIso X.isoSpec.hom.val.base).injective] + erw [Set.preimage_image_eq _ (bijective_of_isIso X.isoSpec.hom.base).injective] · rintro ⟨I, rfl⟩ exact zeroLocus_isClosed X I.carrier diff --git a/Mathlib/AlgebraicGeometry/Cover/Open.lean b/Mathlib/AlgebraicGeometry/Cover/Open.lean index d24109a03e4e4..836bfdf366a33 100644 --- a/Mathlib/AlgebraicGeometry/Cover/Open.lean +++ b/Mathlib/AlgebraicGeometry/Cover/Open.lean @@ -49,7 +49,7 @@ structure OpenCover (X : Scheme.{u}) where /-- given a point of `x : X`, `f x` is the index of the subscheme which contains `x` -/ f : X → J /-- the subschemes covers `X` -/ - covers : ∀ x, x ∈ Set.range (map (f x)).1.base + covers : ∀ x, x ∈ Set.range (map (f x)).base /-- the embedding of subschemes are open immersions -/ IsOpen : ∀ x, IsOpenImmersion (map x) := by infer_instance @@ -65,7 +65,7 @@ def affineCover (X : Scheme.{u}) : OpenCover X where J := X obj x := Spec (X.local_affine x).choose_spec.choose map x := - ((X.local_affine x).choose_spec.choose_spec.some.inv ≫ X.toLocallyRingedSpace.ofRestrict _ : _) + ⟨(X.local_affine x).choose_spec.choose_spec.some.inv ≫ X.toLocallyRingedSpace.ofRestrict _⟩ f x := x covers := by intro x @@ -81,7 +81,7 @@ instance : Inhabited X.OpenCover := ⟨X.affineCover⟩ theorem OpenCover.iUnion_range {X : Scheme.{u}} (𝒰 : X.OpenCover) : - ⋃ i, Set.range (𝒰.map i).1.base = Set.univ := by + ⋃ i, Set.range (𝒰.map i).base = Set.univ := by rw [Set.eq_univ_iff_forall] intro x rw [Set.mem_iUnion] @@ -101,9 +101,9 @@ def OpenCover.bind (f : ∀ x : 𝒰.J, OpenCover (𝒰.obj x)) : OpenCover X wh f x := ⟨_, (f _).f (𝒰.covers x).choose⟩ covers x := by let y := (𝒰.covers x).choose - have hy : (𝒰.map (𝒰.f x)).val.base y = x := (𝒰.covers x).choose_spec + have hy : (𝒰.map (𝒰.f x)).base y = x := (𝒰.covers x).choose_spec rcases (f (𝒰.f x)).covers y with ⟨z, hz⟩ - change x ∈ Set.range ((f (𝒰.f x)).map ((f (𝒰.f x)).f y) ≫ 𝒰.map (𝒰.f x)).1.base + change x ∈ Set.range ((f (𝒰.f x)).map ((f (𝒰.f x)).f y) ≫ 𝒰.map (𝒰.f x)).base use z erw [comp_apply] rw [hz, hy] @@ -133,7 +133,7 @@ def OpenCover.copy {X : Scheme.{u}} (𝒰 : OpenCover X) (J : Type*) (obj : J { J, obj, map f := fun x => e₁.symm (𝒰.f x) covers := fun x => by - rw [e₂, Scheme.comp_val_base, TopCat.coe_comp, Set.range_comp, Set.range_iff_surjective.mpr, + rw [e₂, Scheme.comp_base, TopCat.coe_comp, Set.range_comp, Set.range_iff_surjective.mpr, Set.image_univ, e₁.rightInverse_symm] · exact 𝒰.covers x · rw [← TopCat.epi_iff_surjective]; infer_instance @@ -169,15 +169,15 @@ def OpenCover.pullbackCover {X W : Scheme.{u}} (𝒰 : X.OpenCover) (f : W ⟶ X J := 𝒰.J obj x := pullback f (𝒰.map x) map _ := pullback.fst _ _ - f x := 𝒰.f (f.1.base x) + f x := 𝒰.f (f.base x) covers x := by rw [← - show _ = (pullback.fst _ _ : pullback f (𝒰.map (𝒰.f (f.1.base x))) ⟶ _).1.base from - PreservesPullback.iso_hom_fst Scheme.forgetToTop f (𝒰.map (𝒰.f (f.1.base x)))] + show _ = (pullback.fst _ _ : pullback f (𝒰.map (𝒰.f (f.base x))) ⟶ _).base from + PreservesPullback.iso_hom_fst Scheme.forgetToTop f (𝒰.map (𝒰.f (f.base x)))] -- Porting note: `rw` to `erw` on this single lemma rw [TopCat.coe_comp, Set.range_comp, Set.range_iff_surjective.mpr, Set.image_univ, TopCat.pullback_fst_range] - · obtain ⟨y, h⟩ := 𝒰.covers (f.1.base x) + · obtain ⟨y, h⟩ := 𝒰.covers (f.base x) exact ⟨y, h.symm⟩ · rw [← TopCat.epi_iff_surjective]; infer_instance @@ -199,15 +199,15 @@ def OpenCover.pullbackCover' {X W : Scheme.{u}} (𝒰 : X.OpenCover) (f : W ⟶ J := 𝒰.J obj x := pullback (𝒰.map x) f map _ := pullback.snd _ _ - f x := 𝒰.f (f.1.base x) + f x := 𝒰.f (f.base x) covers x := by rw [← - show _ = (pullback.snd (𝒰.map (𝒰.f (f.1.base x))) f).1.base from - PreservesPullback.iso_hom_snd Scheme.forgetToTop (𝒰.map (𝒰.f (f.1.base x))) f] + show _ = (pullback.snd (𝒰.map (𝒰.f (f.base x))) f).base from + PreservesPullback.iso_hom_snd Scheme.forgetToTop (𝒰.map (𝒰.f (f.base x))) f] -- Porting note: `rw` to `erw` on this single lemma rw [TopCat.coe_comp, Set.range_comp, Set.range_iff_surjective.mpr, Set.image_univ, TopCat.pullback_snd_range] - · obtain ⟨y, h⟩ := 𝒰.covers (f.1.base x) + · obtain ⟨y, h⟩ := 𝒰.covers (f.base x) exact ⟨y, h⟩ · rw [← TopCat.epi_iff_surjective]; infer_instance @@ -217,10 +217,10 @@ def OpenCover.pullbackCover' {X W : Scheme.{u}} (𝒰 : X.OpenCover) (f : W ⟶ def OpenCover.finiteSubcover {X : Scheme.{u}} (𝒰 : OpenCover X) [H : CompactSpace X] : OpenCover X := by have := - @CompactSpace.elim_nhds_subcover _ _ H (fun x : X => Set.range (𝒰.map (𝒰.f x)).1.base) + @CompactSpace.elim_nhds_subcover _ _ H (fun x : X => Set.range (𝒰.map (𝒰.f x)).base) fun x => (IsOpenImmersion.isOpen_range (𝒰.map (𝒰.f x))).mem_nhds (𝒰.covers x) let t := this.choose - have h : ∀ x : X, ∃ y : t, x ∈ Set.range (𝒰.map (𝒰.f y)).1.base := by + have h : ∀ x : X, ∃ y : t, x ∈ Set.range (𝒰.map (𝒰.f y)).base := by intro x have h' : x ∈ (⊤ : Set X) := trivial rw [← Classical.choose_spec this, Set.mem_iUnion] at h' @@ -249,7 +249,7 @@ theorem OpenCover.compactSpace {X : Scheme.{u}} (𝒰 : X.OpenCover) [Finite (asIso (IsOpenImmersion.isoOfRangeEq (𝒰.map i) (X.ofRestrict (Opens.openEmbedding ⟨_, (𝒰.IsOpen i).base_open.isOpen_range⟩)) - Subtype.range_coe.symm).hom.1.base)) + Subtype.range_coe.symm).hom.base)) /-- Given open covers `{ Uᵢ }` and `{ Uⱼ }`, we may form the open cover `{ Uᵢ ∩ Uⱼ }`. -/ def OpenCover.inter {X : Scheme.{u}} (𝒰₁ : Scheme.OpenCover.{v₁} X) @@ -277,7 +277,7 @@ structure AffineOpenCover (X : Scheme.{u}) where /-- given a point of `x : X`, `f x` is the index of the subscheme which contains `x` -/ f : X → J /-- the subschemes covers `X` -/ - covers : ∀ x, x ∈ Set.range (map (f x)).1.base + covers : ∀ x, x ∈ Set.range (map (f x)).base /-- the embedding of subschemes are open immersions -/ IsOpen : ∀ x, IsOpenImmersion (map x) := by infer_instance @@ -462,7 +462,7 @@ def affineBasisCoverOfAffine (R : CommRingCat.{u}) : OpenCover (Spec R) where · exact trivial · -- Porting note: need more hand holding here because Lean knows that -- `CommRing.ofHom ...` is iso, but without `ofHom` Lean does not know what to do - change Epi (Spec.map (CommRingCat.ofHom (algebraMap _ _))).1.base + change Epi (Spec.map (CommRingCat.ofHom (algebraMap _ _))).base infer_instance IsOpen x := AlgebraicGeometry.Scheme.basic_open_isOpenImmersion x @@ -481,8 +481,8 @@ theorem affineBasisCover_obj (X : Scheme.{u}) (i : X.affineBasisCover.J) : theorem affineBasisCover_map_range (X : Scheme.{u}) (x : X) (r : (X.local_affine x).choose_spec.choose) : - Set.range (X.affineBasisCover.map ⟨x, r⟩).1.base = - (X.affineCover.map x).1.base '' (PrimeSpectrum.basicOpen r).1 := by + Set.range (X.affineBasisCover.map ⟨x, r⟩).base = + (X.affineCover.map x).base '' (PrimeSpectrum.basicOpen r).1 := by erw [coe_comp, Set.range_comp] -- Porting note: `congr` fails to see the goal is comparing image of the same function refine congr_arg (_ '' ·) ?_ @@ -491,16 +491,16 @@ theorem affineBasisCover_map_range (X : Scheme.{u}) (x : X) theorem affineBasisCover_is_basis (X : Scheme.{u}) : TopologicalSpace.IsTopologicalBasis {x : Set X | - ∃ a : X.affineBasisCover.J, x = Set.range (X.affineBasisCover.map a).1.base} := by + ∃ a : X.affineBasisCover.J, x = Set.range (X.affineBasisCover.map a).base} := by apply TopologicalSpace.isTopologicalBasis_of_isOpen_of_nhds · rintro _ ⟨a, rfl⟩ exact IsOpenImmersion.isOpen_range (X.affineBasisCover.map a) · rintro a U haU hU rcases X.affineCover.covers a with ⟨x, e⟩ - let U' := (X.affineCover.map (X.affineCover.f a)).1.base ⁻¹' U + let U' := (X.affineCover.map (X.affineCover.f a)).base ⁻¹' U have hxU' : x ∈ U' := by rw [← e] at haU; exact haU rcases PrimeSpectrum.isBasis_basic_opens.exists_subset_of_mem_open hxU' - ((X.affineCover.map (X.affineCover.f a)).1.base.continuous_toFun.isOpen_preimage _ + ((X.affineCover.map (X.affineCover.f a)).base.continuous_toFun.isOpen_preimage _ hU) with ⟨_, ⟨_, ⟨s, rfl⟩, rfl⟩, hxV, hVU⟩ refine ⟨_, ⟨⟨_, s⟩, rfl⟩, ?_, ?_⟩ <;> rw [affineBasisCover_map_range] diff --git a/Mathlib/AlgebraicGeometry/FunctionField.lean b/Mathlib/AlgebraicGeometry/FunctionField.lean index c58bde48fa4db..1b998330cf366 100644 --- a/Mathlib/AlgebraicGeometry/FunctionField.lean +++ b/Mathlib/AlgebraicGeometry/FunctionField.lean @@ -75,9 +75,9 @@ theorem Scheme.germToFunctionField_injective [IsIntegral X] (U : X.Opens) [Nonem theorem genericPoint_eq_of_isOpenImmersion {X Y : Scheme} (f : X ⟶ Y) [H : IsOpenImmersion f] [hX : IrreducibleSpace X] [IrreducibleSpace Y] : - f.1.base (genericPoint X) = genericPoint Y := by + f.base (genericPoint X) = genericPoint Y := by apply ((genericPoint_spec Y).eq _).symm - convert (genericPoint_spec X).image (show Continuous f.1.base by fun_prop) + convert (genericPoint_spec X).image (show Continuous f.base by fun_prop) symm rw [eq_top_iff, Set.top_eq_univ, Set.top_eq_univ] convert subset_closure_inter_of_isPreirreducible_of_isOpen _ H.base_open.isOpen_range _ diff --git a/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean b/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean index 7bd7ef196c89a..bd775fd8d0f99 100644 --- a/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean +++ b/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean @@ -193,9 +193,9 @@ theorem toStalk_stalkMap_toΓSpec (x : X) : exact (X.toΓSpecBase _* X.presheaf).germ_res le_top.hom _ _ /-- The canonical morphism from `X` to the spectrum of its global sections. -/ -@[simps! val_base] +@[simps! base] def toΓSpec : X ⟶ Spec.locallyRingedSpaceObj (Γ.obj (op X)) where - val := X.toΓSpecSheafedSpace + __ := X.toΓSpecSheafedSpace prop := by intro x let p : PrimeSpectrum (Γ.obj (op X)) := X.toΓSpecFun x @@ -220,11 +220,11 @@ def toΓSpec : X ⟶ Spec.locallyRingedSpaceObj (Γ.obj (op X)) where of `Γ(X, ⊤)` under `toΓSpec` agrees with the associated zero locus on `X`. -/ lemma toΓSpec_preimage_zeroLocus_eq {X : LocallyRingedSpace.{u}} (s : Set (X.presheaf.obj (op ⊤))) : - X.toΓSpec.val.base ⁻¹' PrimeSpectrum.zeroLocus s = X.toRingedSpace.zeroLocus s := by + X.toΓSpec.base ⁻¹' PrimeSpectrum.zeroLocus s = X.toRingedSpace.zeroLocus s := by simp only [RingedSpace.zeroLocus] have (i : LocallyRingedSpace.Γ.obj (op X)) (_ : i ∈ s) : ((X.toRingedSpace.basicOpen i).carrier)ᶜ = - X.toΓSpec.val.base ⁻¹' (PrimeSpectrum.basicOpen i).carrierᶜ := by + X.toΓSpec.base ⁻¹' (PrimeSpectrum.basicOpen i).carrierᶜ := by symm erw [Set.preimage_compl, X.toΓSpec_preimage_basicOpen_eq i] erw [Set.iInter₂_congr this] @@ -235,24 +235,24 @@ lemma toΓSpec_preimage_zeroLocus_eq {X : LocallyRingedSpace.{u}} theorem comp_ring_hom_ext {X : LocallyRingedSpace.{u}} {R : CommRingCat.{u}} {f : R ⟶ Γ.obj (op X)} {β : X ⟶ Spec.locallyRingedSpaceObj R} - (w : X.toΓSpec.1.base ≫ (Spec.locallyRingedSpaceMap f).1.base = β.1.base) + (w : X.toΓSpec.base ≫ (Spec.locallyRingedSpaceMap f).base = β.base) (h : ∀ r : R, - f ≫ X.presheaf.map (homOfLE le_top : (Opens.map β.1.base).obj (basicOpen r) ⟶ _).op = - toOpen R (basicOpen r) ≫ β.1.c.app (op (basicOpen r))) : + f ≫ X.presheaf.map (homOfLE le_top : (Opens.map β.base).obj (basicOpen r) ⟶ _).op = + toOpen R (basicOpen r) ≫ β.c.app (op (basicOpen r))) : X.toΓSpec ≫ Spec.locallyRingedSpaceMap f = β := by ext1 -- Porting note: was `apply Spec.basicOpen_hom_ext` refine Spec.basicOpen_hom_ext w ?_ intro r U - rw [LocallyRingedSpace.comp_val_c_app] + rw [LocallyRingedSpace.comp_c_app] erw [toOpen_comp_comap_assoc] rw [Category.assoc] erw [toΓSpecSheafedSpace_app_spec, ← X.presheaf.map_comp] exact h r /-- `toSpecΓ _` is an isomorphism so these are mutually two-sided inverses. -/ -theorem Γ_Spec_left_triangle : toSpecΓ (Γ.obj (op X)) ≫ X.toΓSpec.1.c.app (op ⊤) = 𝟙 _ := by +theorem Γ_Spec_left_triangle : toSpecΓ (Γ.obj (op X)) ≫ X.toΓSpec.c.app (op ⊤) = 𝟙 _ := by unfold toSpecΓ rw [← toOpen_res _ (basicOpen (1 : Γ.obj (op X))) ⊤ (eqToHom basicOpen_one.symm), Category.assoc, NatTrans.naturality, ← Category.assoc] @@ -269,7 +269,7 @@ def identityToΓSpec : 𝟭 LocallyRingedSpace.{u} ⟶ Γ.rightOp ⋙ Spec.toLoc apply LocallyRingedSpace.comp_ring_hom_ext · ext1 x dsimp - show PrimeSpectrum.comap (f.val.c.app (op ⊤)) (X.toΓSpecFun x) = Y.toΓSpecFun (f.val.base x) + show PrimeSpectrum.comap (f.c.app (op ⊤)) (X.toΓSpecFun x) = Y.toΓSpecFun (f.base x) dsimp [toΓSpecFun] -- TODO: this instance was found automatically before #6045 have := @AlgebraicGeometry.LocallyRingedSpace.isLocalRingHomStalkMap X Y @@ -278,14 +278,14 @@ def identityToΓSpec : 𝟭 LocallyRingedSpace.{u} ⟶ Γ.rightOp ⋙ Spec.toLoc congr 2 exact (PresheafedSpace.stalkMap_germ f.1 ⊤ x trivial).symm · intro r - rw [LocallyRingedSpace.comp_val_c_app, ← Category.assoc] - erw [Y.toΓSpecSheafedSpace_app_spec, f.1.c.naturality] + rw [LocallyRingedSpace.comp_c_app, ← Category.assoc] + erw [Y.toΓSpecSheafedSpace_app_spec, f.c.naturality] rfl namespace ΓSpec theorem left_triangle (X : LocallyRingedSpace) : - SpecΓIdentity.inv.app (Γ.obj (op X)) ≫ (identityToΓSpec.app X).val.c.app (op ⊤) = 𝟙 _ := + SpecΓIdentity.inv.app (Γ.obj (op X)) ≫ (identityToΓSpec.app X).c.app (op ⊤) = 𝟙 _ := X.Γ_Spec_left_triangle /-- `SpecΓIdentity` is iso so these are mutually two-sided inverses. -/ @@ -352,7 +352,7 @@ lemma toOpen_comp_locallyRingedSpaceAdjunction_homEquiv_app {X : LocallyRingedSpace} {R : Type u} [CommRing R] (f : Γ.rightOp.obj X ⟶ op (CommRingCat.of R)) (U) : StructureSheaf.toOpen R U.unop ≫ - (locallyRingedSpaceAdjunction.homEquiv X (op <| CommRingCat.of R) f).1.c.app U = + (locallyRingedSpaceAdjunction.homEquiv X (op <| CommRingCat.of R) f).c.app U = f.unop ≫ X.presheaf.map (homOfLE le_top).op := by rw [← StructureSheaf.toOpen_res _ _ _ (homOfLE le_top), Category.assoc, NatTrans.naturality _ (homOfLE (le_top (a := U.unop))).op, @@ -364,27 +364,25 @@ lemma toOpen_comp_locallyRingedSpaceAdjunction_homEquiv_app rfl /-- The adjunction `Γ ⊣ Spec` from `CommRingᵒᵖ` to `Scheme`. -/ -def adjunction : Scheme.Γ.rightOp ⊣ Scheme.Spec.{u} := - Adjunction.mk' { - homEquiv := fun X Y ↦ locallyRingedSpaceAdjunction.{u}.homEquiv X.toLocallyRingedSpace Y - unit := - { app := fun X ↦ locallyRingedSpaceAdjunction.{u}.unit.app X.toLocallyRingedSpace - naturality := fun _ _ f ↦ locallyRingedSpaceAdjunction.{u}.unit.naturality f } - counit := (NatIso.op Scheme.SpecΓIdentity.{u}).inv - homEquiv_unit := rfl - homEquiv_counit := rfl } +def adjunction : Scheme.Γ.rightOp ⊣ Scheme.Spec.{u} where + unit := + { app := fun X ↦ ⟨locallyRingedSpaceAdjunction.{u}.unit.app X.toLocallyRingedSpace⟩ + naturality := fun _ _ f ↦ + Scheme.Hom.ext' (locallyRingedSpaceAdjunction.{u}.unit.naturality f.toLRSHom) } + counit := (NatIso.op Scheme.SpecΓIdentity.{u}).inv + left_triangle_components Y := + locallyRingedSpaceAdjunction.left_triangle_components Y.toLocallyRingedSpace + right_triangle_components R := + Scheme.Hom.ext' <| locallyRingedSpaceAdjunction.right_triangle_components R theorem adjunction_homEquiv_apply {X : Scheme} {R : CommRingCatᵒᵖ} (f : (op <| Scheme.Γ.obj <| op X) ⟶ R) : - ΓSpec.adjunction.homEquiv X R f = locallyRingedSpaceAdjunction.homEquiv X.1 R f := rfl - -theorem adjunction_homEquiv (X : Scheme) (R : CommRingCatᵒᵖ) : - ΓSpec.adjunction.homEquiv X R = locallyRingedSpaceAdjunction.homEquiv X.1 R := rfl + ΓSpec.adjunction.homEquiv X R f = ⟨locallyRingedSpaceAdjunction.homEquiv X.1 R f⟩ := rfl theorem adjunction_homEquiv_symm_apply {X : Scheme} {R : CommRingCatᵒᵖ} (f : X ⟶ Scheme.Spec.obj R) : (ΓSpec.adjunction.homEquiv X R).symm f = - (locallyRingedSpaceAdjunction.homEquiv X.1 R).symm f := rfl + (locallyRingedSpaceAdjunction.homEquiv X.1 R).symm f.toLRSHom := rfl theorem adjunction_counit_app' {R : CommRingCatᵒᵖ} : ΓSpec.adjunction.counit.app R = locallyRingedSpaceAdjunction.counit.app R := rfl @@ -413,9 +411,9 @@ instance isIso_adjunction_counit : IsIso ΓSpec.adjunction.counit := by end ΓSpec -theorem Scheme.toSpecΓ_val_base (X : Scheme.{u}) (x) : - (Scheme.toSpecΓ X).1.base x = - (Spec.map (X.presheaf.germ ⊤ x trivial)).1.base (LocalRing.closedPoint _) := rfl +theorem Scheme.toSpecΓ_base (X : Scheme.{u}) (x) : + (Scheme.toSpecΓ X).base x = + (Spec.map (X.presheaf.germ ⊤ x trivial)).base (LocalRing.closedPoint _) := rfl @[reassoc (attr := simp)] theorem Scheme.toSpecΓ_naturality {X Y : Scheme.{u}} (f : X ⟶ Y) : diff --git a/Mathlib/AlgebraicGeometry/Gluing.lean b/Mathlib/AlgebraicGeometry/Gluing.lean index 2907eed57c7e5..9eeee3d9e71e3 100644 --- a/Mathlib/AlgebraicGeometry/Gluing.lean +++ b/Mathlib/AlgebraicGeometry/Gluing.lean @@ -123,9 +123,8 @@ def gluedScheme : Scheme := by D.toLocallyRingedSpaceGlueData.toGlueData.glued intro x obtain ⟨i, y, rfl⟩ := D.toLocallyRingedSpaceGlueData.ι_jointly_surjective x - refine ⟨?_, ?_ ≫ D.toLocallyRingedSpaceGlueData.toGlueData.ι i, ?_⟩ - swap - · exact (D.U i).affineCover.map y + refine ⟨_, ((D.U i).affineCover.map y).toLRSHom ≫ + D.toLocallyRingedSpaceGlueData.toGlueData.ι i, ?_⟩ constructor · erw [TopCat.coe_comp, Set.range_comp] -- now `erw` after #13170 refine Set.mem_image_of_mem _ ?_ @@ -157,14 +156,15 @@ abbrev isoLocallyRingedSpace : 𝖣.gluedIso forgetToLocallyRingedSpace theorem ι_isoLocallyRingedSpace_inv (i : D.J) : - D.toLocallyRingedSpaceGlueData.toGlueData.ι i ≫ D.isoLocallyRingedSpace.inv = 𝖣.ι i := + D.toLocallyRingedSpaceGlueData.toGlueData.ι i ≫ + D.isoLocallyRingedSpace.inv = (𝖣.ι i).toLRSHom := 𝖣.ι_gluedIso_inv forgetToLocallyRingedSpace i instance ι_isOpenImmersion (i : D.J) : IsOpenImmersion (𝖣.ι i) := by - rw [← D.ι_isoLocallyRingedSpace_inv]; infer_instance + rw [IsOpenImmersion, ← D.ι_isoLocallyRingedSpace_inv]; infer_instance theorem ι_jointly_surjective (x : 𝖣.glued.carrier) : - ∃ (i : D.J) (y : (D.U i).carrier), (D.ι i).1.base y = x := + ∃ (i : D.J) (y : (D.U i).carrier), (D.ι i).base y = x := 𝖣.ι_jointly_surjective (forgetToTop ⋙ forget TopCat) x -- Porting note: promote to higher priority to short circuit simplifier @@ -206,24 +206,24 @@ def isoCarrier : @[simp] theorem ι_isoCarrier_inv (i : D.J) : - (D_).ι i ≫ D.isoCarrier.inv = (D.ι i).1.base := by + (D_).ι i ≫ D.isoCarrier.inv = (D.ι i).base := by delta isoCarrier rw [Iso.trans_inv, GlueData.ι_gluedIso_inv_assoc, Functor.mapIso_inv, Iso.trans_inv, Functor.mapIso_inv, Iso.trans_inv, SheafedSpace.forgetToPresheafedSpace_map, forget_map, - forget_map, ← comp_base, ← Category.assoc, + forget_map, ← PresheafedSpace.comp_base, ← Category.assoc, D.toLocallyRingedSpaceGlueData.toSheafedSpaceGlueData.ι_isoPresheafedSpace_inv i] erw [← Category.assoc, D.toLocallyRingedSpaceGlueData.ι_isoSheafedSpace_inv i] - change (_ ≫ D.isoLocallyRingedSpace.inv).1.base = _ + change (_ ≫ D.isoLocallyRingedSpace.inv).base = _ rw [D.ι_isoLocallyRingedSpace_inv i] /-- An equivalence relation on `Σ i, D.U i` that holds iff `𝖣 .ι i x = 𝖣 .ι j y`. See `AlgebraicGeometry.Scheme.GlueData.ι_eq_iff`. -/ def Rel (a b : Σ i, ((D.U i).carrier : Type _)) : Prop := a = b ∨ - ∃ x : (D.V (a.1, b.1)).carrier, (D.f _ _).1.base x = a.2 ∧ (D.t _ _ ≫ D.f _ _).1.base x = b.2 + ∃ x : (D.V (a.1, b.1)).carrier, (D.f _ _).base x = a.2 ∧ (D.t _ _ ≫ D.f _ _).base x = b.2 theorem ι_eq_iff (i j : D.J) (x : (D.U i).carrier) (y : (D.U j).carrier) : - (𝖣.ι i).1.base x = (𝖣.ι j).1.base y ↔ D.Rel ⟨i, x⟩ ⟨j, y⟩ := by + (𝖣.ι i).base x = (𝖣.ι j).base y ↔ D.Rel ⟨i, x⟩ ⟨j, y⟩ := by refine Iff.trans ?_ (TopCat.GlueData.ι_eq_iff_rel D.toLocallyRingedSpaceGlueData.toSheafedSpaceGlueData.toPresheafedSpaceGlueData.toTopGlueData @@ -233,7 +233,7 @@ theorem ι_eq_iff (i j : D.J) (x : (D.U i).carrier) (y : (D.U j).carrier) : rfl -- `rfl` was not needed before #13170 · infer_instance -theorem isOpen_iff (U : Set D.glued.carrier) : IsOpen U ↔ ∀ i, IsOpen ((D.ι i).1.base ⁻¹' U) := by +theorem isOpen_iff (U : Set D.glued.carrier) : IsOpen U ↔ ∀ i, IsOpen ((D.ι i).base ⁻¹' U) := by rw [← (TopCat.homeoOfIso D.isoCarrier.symm).isOpen_preimage, TopCat.GlueData.isOpen_iff] apply forall_congr' intro i @@ -336,13 +336,13 @@ def fromGlued : 𝒰.gluedCover.glued ⟶ X := by theorem ι_fromGlued (x : 𝒰.J) : 𝒰.gluedCover.ι x ≫ 𝒰.fromGlued = 𝒰.map x := Multicoequalizer.π_desc _ _ _ _ _ -theorem fromGlued_injective : Function.Injective 𝒰.fromGlued.1.base := by +theorem fromGlued_injective : Function.Injective 𝒰.fromGlued.base := by intro x y h obtain ⟨i, x, rfl⟩ := 𝒰.gluedCover.ι_jointly_surjective x obtain ⟨j, y, rfl⟩ := 𝒰.gluedCover.ι_jointly_surjective y rw [← comp_apply, ← comp_apply] at h - simp_rw [← SheafedSpace.comp_base, ← LocallyRingedSpace.comp_val] at h - erw [ι_fromGlued, ι_fromGlued] at h + simp_rw [← Scheme.comp_base] at h + rw [ι_fromGlued, ι_fromGlued] at h let e := (TopCat.pullbackConeIsLimit _ _).conePointUniqueUpToIso (isLimitOfHasPullbackOfPreservesLimit Scheme.forgetToTop (𝒰.map i) (𝒰.map j)) @@ -363,12 +363,12 @@ instance fromGlued_stalk_iso (x : 𝒰.gluedCover.glued.carrier) : rw [this] infer_instance -theorem fromGlued_open_map : IsOpenMap 𝒰.fromGlued.1.base := by +theorem fromGlued_open_map : IsOpenMap 𝒰.fromGlued.base := by intro U hU rw [isOpen_iff_forall_mem_open] intro x hx rw [𝒰.gluedCover.isOpen_iff] at hU - use 𝒰.fromGlued.val.base '' U ∩ Set.range (𝒰.map (𝒰.f x)).1.base + use 𝒰.fromGlued.base '' U ∩ Set.range (𝒰.map (𝒰.f x)).base use Set.inter_subset_left constructor · rw [← Set.image_preimage_eq_inter_range] @@ -379,15 +379,15 @@ theorem fromGlued_open_map : IsOpenMap 𝒰.fromGlued.1.base := by exact Set.preimage_image_eq _ 𝒰.fromGlued_injective · exact ⟨hx, 𝒰.covers x⟩ -theorem fromGlued_openEmbedding : OpenEmbedding 𝒰.fromGlued.1.base := +theorem fromGlued_openEmbedding : OpenEmbedding 𝒰.fromGlued.base := openEmbedding_of_continuous_injective_open (by fun_prop) 𝒰.fromGlued_injective 𝒰.fromGlued_open_map -instance : Epi 𝒰.fromGlued.val.base := by +instance : Epi 𝒰.fromGlued.base := by rw [TopCat.epi_iff_surjective] intro x obtain ⟨y, h⟩ := 𝒰.covers x - use (𝒰.gluedCover.ι (𝒰.f x)).1.base y + use (𝒰.gluedCover.ι (𝒰.f x)).base y rw [← comp_apply] rw [← 𝒰.ι_fromGlued (𝒰.f x)] at h exact h @@ -399,7 +399,7 @@ instance : IsIso 𝒰.fromGlued := let F := Scheme.forgetToLocallyRingedSpace ⋙ LocallyRingedSpace.forgetToSheafedSpace ⋙ SheafedSpace.forgetToPresheafedSpace have : IsIso (F.map (fromGlued 𝒰)) := by - change @IsIso (PresheafedSpace _) _ _ _ 𝒰.fromGlued.val + change IsIso 𝒰.fromGlued.toPshHom apply PresheafedSpace.IsOpenImmersion.to_iso isIso_of_reflects_iso _ F diff --git a/Mathlib/AlgebraicGeometry/Limits.lean b/Mathlib/AlgebraicGeometry/Limits.lean index 5124bdbd6e750..d5614273639d9 100644 --- a/Mathlib/AlgebraicGeometry/Limits.lean +++ b/Mathlib/AlgebraicGeometry/Limits.lean @@ -58,8 +58,7 @@ def Scheme.emptyTo (X : Scheme.{u}) : ∅ ⟶ X := @[ext] theorem Scheme.empty_ext {X : Scheme.{u}} (f g : ∅ ⟶ X) : f = g := -- Porting note (#11041): `ext` regression - LocallyRingedSpace.Hom.ext <| PresheafedSpace.ext _ _ (by ext a; exact PEmpty.elim a) <| - NatTrans.ext <| funext fun a => by aesop_cat + Scheme.Hom.ext' (Subsingleton.elim (α := ∅ ⟶ _) _ _) theorem Scheme.eq_emptyTo {X : Scheme.{u}} (f : ∅ ⟶ X) : f = Scheme.emptyTo X := Scheme.empty_ext f (Scheme.emptyTo X) @@ -89,8 +88,8 @@ instance (priority := 100) isOpenImmersion_of_isEmpty {X Y : Scheme} (f : X ⟶ instance (priority := 100) isIso_of_isEmpty {X Y : Scheme} (f : X ⟶ Y) [IsEmpty Y] : IsIso f := by - haveI : IsEmpty X := f.1.base.1.isEmpty - have : Epi f.1.base := by + haveI : IsEmpty X := f.base.1.isEmpty + have : Epi f.base := by rw [TopCat.epi_iff_surjective]; rintro (x : Y) exact isEmptyElim x apply IsOpenImmersion.to_iso @@ -110,7 +109,7 @@ instance : HasInitial Scheme.{u} := hasInitial_of_unique ∅ instance initial_isEmpty : IsEmpty (⊥_ Scheme) := - ⟨fun x => ((initial.to Scheme.empty : _).1.base x).elim⟩ + ⟨fun x => ((initial.to Scheme.empty : _).base x).elim⟩ theorem isAffineOpen_bot (X : Scheme) : IsAffineOpen (⊥ : X.Opens) := @isAffine_of_isEmpty _ (inferInstanceAs (IsEmpty (∅ : Set X))) @@ -242,7 +241,7 @@ instance (i) : IsOpenImmersion (Sigma.ι f i) := by infer_instance lemma sigmaι_eq_iff (i j : ι) (x y) : - (Sigma.ι f i).1.base x = (Sigma.ι f j).1.base y ↔ + (Sigma.ι f i).base x = (Sigma.ι f j).base y ↔ (Sigma.mk i x : Σ i, f i) = Sigma.mk j y := by constructor · intro H @@ -269,10 +268,10 @@ lemma disjoint_opensRange_sigmaι (i j : ι) (h : i ≠ j) : obtain ⟨rfl⟩ := (sigmaι_eq_iff _ _ _ _ _).mp hy cases h rfl -lemma exists_sigmaι_eq (x : (∐ f : _)) : ∃ i y, (Sigma.ι f i).1.base y = x := by - obtain ⟨i, y, e⟩ := (disjointGlueData f).ι_jointly_surjective ((sigmaIsoGlued f).hom.1.base x) +lemma exists_sigmaι_eq (x : (∐ f : _)) : ∃ i y, (Sigma.ι f i).base y = x := by + obtain ⟨i, y, e⟩ := (disjointGlueData f).ι_jointly_surjective ((sigmaIsoGlued f).hom.base x) refine ⟨i, y, (sigmaIsoGlued f).hom.openEmbedding.inj ?_⟩ - rwa [← Scheme.comp_val_base_apply, ι_sigmaIsoGlued_hom] + rwa [← Scheme.comp_base_apply, ι_sigmaIsoGlued_hom] lemma iSup_opensRange_sigmaι : ⨆ i, (Sigma.ι f i).opensRange = ⊤ := eq_top_iff.mpr fun x ↦ by simpa using exists_sigmaι_eq f x @@ -295,7 +294,7 @@ def sigmaMk : (Σ i, f i) ≃ₜ (∐ f : _) := @[simp] lemma sigmaMk_mk (i) (x : f i) : - sigmaMk f (.mk i x) = (Sigma.ι f i).1.base x := by + sigmaMk f (.mk i x) = (Sigma.ι f i).base x := by show ((TopCat.sigmaCofan (fun x ↦ (f x).toTopCat)).inj i ≫ (colimit.isoColimitCocone ⟨_, TopCat.sigmaCofanIsColimit _⟩).inv ≫ _) x = Scheme.forgetToTop.map (Sigma.ι f i) x @@ -324,8 +323,8 @@ instance : IsOpenImmersion (coprod.inr : Y ⟶ X ⨿ Y) := by rw [← ι_right_coprodIsoSigma_inv]; infer_instance lemma isCompl_range_inl_inr : - IsCompl (Set.range (coprod.inl : X ⟶ X ⨿ Y).1.base) - (Set.range (coprod.inr : Y ⟶ X ⨿ Y).1.base) := + IsCompl (Set.range (coprod.inl : X ⟶ X ⨿ Y).base) + (Set.range (coprod.inr : Y ⟶ X ⨿ Y).base) := ((TopCat.binaryCofan_isColimit_iff _).mp ⟨mapIsColimitOfPreservesOfIsColimit Scheme.forgetToTop _ _ (coprodIsCoprod X Y)⟩).2.2 @@ -343,7 +342,7 @@ def coprodMk : X ⊕ Y ≃ₜ (X ⨿ Y : Scheme.{u}) := @[simp] lemma coprodMk_inl (x : X) : - coprodMk X Y (.inl x) = (coprod.inl : X ⟶ X ⨿ Y).1.base x := by + coprodMk X Y (.inl x) = (coprod.inl : X ⟶ X ⨿ Y).base x := by show ((TopCat.binaryCofan X Y).inl ≫ (colimit.isoColimitCocone ⟨_, TopCat.binaryCofanIsColimit _ _⟩).inv ≫ _) x = Scheme.forgetToTop.map coprod.inl x @@ -353,7 +352,7 @@ lemma coprodMk_inl (x : X) : @[simp] lemma coprodMk_inr (x : Y) : - coprodMk X Y (.inr x) = (coprod.inr : Y ⟶ X ⨿ Y).1.base x := by + coprodMk X Y (.inr x) = (coprod.inr : Y ⟶ X ⨿ Y).base x := by show ((TopCat.binaryCofan X Y).inr ≫ (colimit.isoColimitCocone ⟨_, TopCat.binaryCofanIsColimit _ _⟩).inv ≫ _) x = Scheme.forgetToTop.map coprod.inr x @@ -397,10 +396,10 @@ lemma coprodSpec_inr : coprod.inr ≫ coprodSpec R S = coprod.inr_desc _ _ lemma coprodSpec_coprodMk (x) : - (coprodSpec R S).1.base (coprodMk _ _ x) = (PrimeSpectrum.primeSpectrumProd R S).symm x := by + (coprodSpec R S).base (coprodMk _ _ x) = (PrimeSpectrum.primeSpectrumProd R S).symm x := by apply PrimeSpectrum.ext obtain (x | x) := x <;> - simp only [coprodMk_inl, coprodMk_inr, ← Scheme.comp_val_base_apply, + simp only [coprodMk_inl, coprodMk_inr, ← Scheme.comp_base_apply, coprodSpec, coprod.inl_desc, coprod.inr_desc] · show Ideal.comap _ _ = x.asIdeal.prod ⊤ ext; simp [Ideal.prod, CommRingCat.ofHom] @@ -408,7 +407,7 @@ lemma coprodSpec_coprodMk (x) : ext; simp [Ideal.prod, CommRingCat.ofHom] lemma coprodSpec_apply (x) : - (coprodSpec R S).1.base x = (PrimeSpectrum.primeSpectrumProd R S).symm + (coprodSpec R S).base x = (PrimeSpectrum.primeSpectrumProd R S).symm ((coprodMk (Spec (.of R)) (Spec (.of S))).symm x) := by rw [← coprodSpec_coprodMk, Homeomorph.apply_symm_apply] @@ -467,7 +466,7 @@ noncomputable instance : PreservesColimitsOfShape (Discrete PEmpty.{1}) Scheme.Spec := by have : IsEmpty (Scheme.Spec.obj (⊥_ CommRingCatᵒᵖ)) := @Function.isEmpty _ _ spec_punit_isEmpty (Scheme.Spec.mapIso - (initialIsoIsInitial (initialOpOfTerminal CommRingCat.punitIsTerminal))).hom.1.base + (initialIsoIsInitial (initialOpOfTerminal CommRingCat.punitIsTerminal))).hom.base have := preservesInitialOfIso Scheme.Spec (asIso (initial.to _)) exact preservesColimitsOfShapePemptyOfPreservesInitial _ diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Affine.lean b/Mathlib/AlgebraicGeometry/Morphisms/Affine.lean index c990b0ee818d7..ecd431a4735b4 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Affine.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Affine.lean @@ -66,7 +66,7 @@ instance (priority := 900) [IsAffineHom f] : QuasiCompact f := instance [IsAffineHom f] [IsAffineHom g] : IsAffineHom (f ≫ g) := by constructor intros U hU - rw [Scheme.comp_val_base, Opens.map_comp_obj] + rw [Scheme.comp_base, Opens.map_comp_obj] apply IsAffineHom.isAffine_preimage apply IsAffineHom.isAffine_preimage exact hU @@ -158,7 +158,7 @@ instance : HasAffineProperty @IsAffineHom fun X _ _ _ ↦ IsAffine X where rw [Scheme.preimage_basicOpen] exact (isAffineOpen_top X).basicOpen _ · intro X Y _ f S hS hS' - apply_fun Ideal.map (f.1.c.app (op ⊤)) at hS + apply_fun Ideal.map (f.app ⊤) at hS rw [Ideal.map_span, Ideal.map_top] at hS apply isAffine_of_isAffineOpen_basicOpen _ hS have : ∀ i : S, IsAffineOpen (f⁻¹ᵁ Y.basicOpen i.1) := hS' diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Basic.lean b/Mathlib/AlgebraicGeometry/Morphisms/Basic.lean index a9f5d34528df4..50efd7b087298 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Basic.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Basic.lean @@ -437,7 +437,7 @@ instance (P : AffineTargetMorphismProperty) [P.toProperty.RespectsIso] : rintro ⟨U, hU : IsAffineOpen U⟩; dsimp haveI : IsAffine _ := hU.preimage_of_isIso e.hom rw [morphismRestrict_comp, P.cancel_right_of_respectsIso] - exact H ⟨(Opens.map e.hom.val.base).obj U, hU.preimage_of_isIso e.hom⟩ + exact H ⟨(Opens.map e.hom.base).obj U, hU.preimage_of_isIso e.hom⟩ /-- `HasAffineProperty P Q` is a type class asserting that `P` is local at the target, and over affine diff --git a/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean b/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean index e0f201805fa3b..592f5829e66fa 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean @@ -38,13 +38,13 @@ namespace AlgebraicGeometry topological map is a closed embedding and the induced stalk maps are surjective. -/ @[mk_iff] class IsClosedImmersion {X Y : Scheme} (f : X ⟶ Y) : Prop where - base_closed : ClosedEmbedding f.1.base + base_closed : ClosedEmbedding f.base surj_on_stalks : ∀ x, Function.Surjective (f.stalkMap x) namespace IsClosedImmersion lemma closedEmbedding {X Y : Scheme} (f : X ⟶ Y) - [IsClosedImmersion f] : ClosedEmbedding f.1.base := + [IsClosedImmersion f] : ClosedEmbedding f.base := IsClosedImmersion.base_closed lemma eq_inf : @IsClosedImmersion = (topologically ClosedEmbedding) ⊓ @@ -54,12 +54,12 @@ lemma eq_inf : @IsClosedImmersion = (topologically ClosedEmbedding) ⊓ rfl lemma iff_isPreimmersion {X Y : Scheme} {f : X ⟶ Y} : - IsClosedImmersion f ↔ IsPreimmersion f ∧ IsClosed (Set.range f.1.base) := by + IsClosedImmersion f ↔ IsPreimmersion f ∧ IsClosed (Set.range f.base) := by rw [and_comm, isClosedImmersion_iff, isPreimmersion_iff, ← and_assoc, closedEmbedding_iff, @and_comm (Embedding _)] lemma of_isPreimmersion {X Y : Scheme} (f : X ⟶ Y) [IsPreimmersion f] - (hf : IsClosed (Set.range f.1.base)) : IsClosedImmersion f := + (hf : IsClosed (Set.range f.base)) : IsClosedImmersion f := iff_isPreimmersion.mpr ⟨‹_›, hf⟩ instance (priority := 900) {X Y : Scheme} (f : X ⟶ Y) [IsClosedImmersion f] : IsPreimmersion f := @@ -67,7 +67,7 @@ instance (priority := 900) {X Y : Scheme} (f : X ⟶ Y) [IsClosedImmersion f] : /-- Isomorphisms are closed immersions. -/ instance {X Y : Scheme} (f : X ⟶ Y) [IsIso f] : IsClosedImmersion f where - base_closed := Homeomorph.closedEmbedding <| TopCat.homeoOfIso (asIso f.1.base) + base_closed := Homeomorph.closedEmbedding <| TopCat.homeoOfIso (asIso f.base) surj_on_stalks := fun _ ↦ (ConcreteCategory.bijective_of_isIso _).2 instance : MorphismProperty.IsMultiplicative @IsClosedImmersion where @@ -75,7 +75,7 @@ instance : MorphismProperty.IsMultiplicative @IsClosedImmersion where comp_mem {X Y Z} f g hf hg := by refine ⟨hg.base_closed.comp hf.base_closed, fun x ↦ ?_⟩ rw [Scheme.stalkMap_comp] - exact (hf.surj_on_stalks x).comp (hg.surj_on_stalks (f.1.1 x)) + exact (hf.surj_on_stalks x).comp (hg.surj_on_stalks (f.base x)) /-- Composition of closed immersions is a closed immersion. -/ instance comp {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) [IsClosedImmersion f] @@ -120,7 +120,7 @@ theorem of_comp {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) [IsClosedImmersion [IsClosedImmersion (f ≫ g)] : IsClosedImmersion f where base_closed := by have h := closedEmbedding (f ≫ g) - rw [Scheme.comp_val_base] at h + rw [Scheme.comp_base] at h apply closedEmbedding_of_continuous_injective_closed (Scheme.Hom.continuous f) · exact Function.Injective.of_comp h.inj · intro Z hZ @@ -129,7 +129,7 @@ theorem of_comp {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) [IsClosedImmersion exact ClosedEmbedding.isClosedMap h _ hZ surj_on_stalks x := by have h := (f ≫ g).stalkMap_surjective x - simp_rw [Scheme.comp_val, Scheme.stalkMap_comp] at h + simp_rw [Scheme.stalkMap_comp] at h exact Function.Surjective.of_comp h instance {X Y : Scheme} (f : X ⟶ Y) [IsClosedImmersion f] : QuasiCompact f where diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Constructors.lean b/Mathlib/AlgebraicGeometry/Morphisms/Constructors.lean index d817f2b6c1a4c..6cfb3141006be 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Constructors.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Constructors.lean @@ -165,7 +165,7 @@ theorem universally_isLocalAtTarget (P : MorphismProperty Scheme) apply hP₂ _ (fun i ↦ i₂ ⁻¹ᵁ U i) · rw [← top_le_iff] at hU ⊢ rintro x - - simpa using @hU (i₂.1.base x) trivial + simpa using @hU (i₂.base x) trivial · rintro i refine H _ ((X'.restrictIsoOfEq ?_).hom ≫ i₁ ∣_ _) (i₂ ∣_ _) _ ?_ · exact congr($(h.1.1) ⁻¹ᵁ U i) @@ -184,7 +184,7 @@ section Topologically /-- `topologically P` holds for a morphism if the underlying topological map satisfies `P`. -/ def topologically (P : ∀ {α β : Type u} [TopologicalSpace α] [TopologicalSpace β] (_ : α → β), Prop) : - MorphismProperty Scheme.{u} := fun _ _ f => P f.1.base + MorphismProperty Scheme.{u} := fun _ _ f => P f.base variable (P : ∀ {α β : Type u} [TopologicalSpace α] [TopologicalSpace β] (_ : α → β), Prop) @@ -205,7 +205,7 @@ lemma topologically_iso_le MorphismProperty.isomorphisms Scheme ≤ (topologically P) := by intro X Y e (he : IsIso e) have : IsIso e := he - exact hP (TopCat.homeoOfIso (asIso e.val.base)) + exact hP (TopCat.homeoOfIso (asIso e.base)) /-- If a property of maps of topological spaces is satisfied by homeomorphisms and is stable under composition, the induced property on schemes respects isomorphisms. -/ @@ -230,11 +230,11 @@ lemma topologically_isLocalAtTarget IsLocalAtTarget (topologically P) := by apply IsLocalAtTarget.mk' · intro X Y f U hf - simp_rw [topologically, morphismRestrict_val_base] - exact hP₂ f.val.base U.carrier hf + simp_rw [topologically, morphismRestrict_base] + exact hP₂ f.base U.carrier hf · intro X Y f ι U hU hf - apply hP₃ f.val.base U hU f.val.base.continuous fun i ↦ ?_ - rw [← morphismRestrict_val_base] + apply hP₃ f.base U hU f.base.continuous fun i ↦ ?_ + rw [← morphismRestrict_base] exact hf i end Topologically @@ -255,8 +255,8 @@ lemma stalkwise_respectsIso (hP : RingHom.RespectsIso P) : simp only [stalkwise, Scheme.comp_coeBase, TopCat.coe_comp, Function.comp_apply] intro x rw [Scheme.stalkMap_comp] - exact (RingHom.RespectsIso.cancel_right_isIso hP _ _).mpr <| hf (e.val.base x) - postcomp {X Y Z} e (he : IsIso e) f hf := by + exact (RingHom.RespectsIso.cancel_right_isIso hP _ _).mpr <| hf (e.base x) + postcomp {X Y Z} e (he : IsIso _) f hf := by simp only [stalkwise, Scheme.comp_coeBase, TopCat.coe_comp, Function.comp_apply] intro x rw [Scheme.stalkMap_comp] @@ -273,7 +273,7 @@ lemma stalkwiseIsLocalAtTarget_of_respectsIso (hP : RingHom.RespectsIso P) : apply ((RingHom.toMorphismProperty P).arrow_mk_iso_iff <| morphismRestrictStalkMap f U x).mpr <| hf _ · intro X Y f ι U hU hf x - have hy : f.val.base x ∈ iSup U := by rw [hU]; trivial + have hy : f.base x ∈ iSup U := by rw [hU]; trivial obtain ⟨i, hi⟩ := Opens.mem_iSup.mp hy exact ((RingHom.toMorphismProperty P).arrow_mk_iso_iff <| morphismRestrictStalkMap f (U i) ⟨x, hi⟩).mp <| hf i ⟨x, hi⟩ diff --git a/Mathlib/AlgebraicGeometry/Morphisms/IsIso.lean b/Mathlib/AlgebraicGeometry/Morphisms/IsIso.lean index ad9b6dcc7c5f6..384f3493f9a1e 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/IsIso.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/IsIso.lean @@ -30,8 +30,8 @@ lemma isomorphisms_eq_stalkwise : congr 1 ext X Y f exact ⟨fun H ↦ inferInstanceAs (IsIso (TopCat.isoOfHomeo - (H.1.1.toHomeomeomorph_of_surjective H.2)).hom), fun (_ : IsIso f.1.base) ↦ - let e := (TopCat.homeoOfIso <| asIso f.1.base); ⟨e.openEmbedding, e.surjective⟩⟩ + (H.1.1.toHomeomeomorph_of_surjective H.2)).hom), fun (_ : IsIso f.base) ↦ + let e := (TopCat.homeoOfIso <| asIso f.base); ⟨e.openEmbedding, e.surjective⟩⟩ instance : IsLocalAtTarget (isomorphisms Scheme) := isomorphisms_eq_isOpenImmersion_inf_surjective ▸ inferInstance diff --git a/Mathlib/AlgebraicGeometry/Morphisms/OpenImmersion.lean b/Mathlib/AlgebraicGeometry/Morphisms/OpenImmersion.lean index 3fc310d18e207..43a94b8d2d1a0 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/OpenImmersion.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/OpenImmersion.lean @@ -29,7 +29,7 @@ namespace AlgebraicGeometry variable {X Y Z : Scheme.{u}} (f : X ⟶ Y) (g : Y ⟶ Z) theorem isOpenImmersion_iff_stalk {f : X ⟶ Y} : IsOpenImmersion f ↔ - OpenEmbedding f.1.base ∧ ∀ x, IsIso (f.stalkMap x) := by + OpenEmbedding f.base ∧ ∀ x, IsIso (f.stalkMap x) := by constructor · intro h; exact ⟨h.1, inferInstance⟩ · rintro ⟨h₁, h₂⟩; exact IsOpenImmersion.of_stalk_iso f h₁ @@ -43,7 +43,7 @@ theorem isOpenImmersion_eq_inf : instance isOpenImmersion_isStableUnderComposition : MorphismProperty.IsStableUnderComposition @IsOpenImmersion where - comp_mem f g _ _ := LocallyRingedSpace.IsOpenImmersion.comp f g + comp_mem f g _ _ := LocallyRingedSpace.IsOpenImmersion.comp f.toLRSHom g.toLRSHom instance isOpenImmersion_respectsIso : MorphismProperty.RespectsIso @IsOpenImmersion := by apply MorphismProperty.respectsIso_of_isStableUnderComposition diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Preimmersion.lean b/Mathlib/AlgebraicGeometry/Morphisms/Preimmersion.lean index b9472fc3ea550..a5fd6903836cb 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Preimmersion.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Preimmersion.lean @@ -33,10 +33,10 @@ namespace AlgebraicGeometry topological spaces is an embedding and the induced morphisms of stalks are all surjective. -/ @[mk_iff] class IsPreimmersion {X Y : Scheme} (f : X ⟶ Y) : Prop where - base_embedding : Embedding f.1.base + base_embedding : Embedding f.base surj_on_stalks : ∀ x, Function.Surjective (f.stalkMap x) -lemma Scheme.Hom.embedding {X Y : Scheme} (f : Hom X Y) [IsPreimmersion f] : Embedding f.1.base := +lemma Scheme.Hom.embedding {X Y : Scheme} (f : Hom X Y) [IsPreimmersion f] : Embedding f.base := IsPreimmersion.base_embedding lemma Scheme.Hom.stalkMap_surjective {X Y : Scheme} (f : Hom X Y) [IsPreimmersion f] (x) : @@ -68,7 +68,7 @@ instance : MorphismProperty.IsMultiplicative @IsPreimmersion where comp_mem {X Y Z} f g hf hg := by refine ⟨hg.base_embedding.comp hf.base_embedding, fun x ↦ ?_⟩ rw [Scheme.stalkMap_comp] - exact (hf.surj_on_stalks x).comp (hg.surj_on_stalks (f.1.1 x)) + exact (hf.surj_on_stalks x).comp (hg.surj_on_stalks (f.base x)) instance comp {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) [IsPreimmersion f] [IsPreimmersion g] : IsPreimmersion (f ≫ g) := diff --git a/Mathlib/AlgebraicGeometry/Morphisms/QuasiCompact.lean b/Mathlib/AlgebraicGeometry/Morphisms/QuasiCompact.lean index cf3a6124cc51f..2e50b205e0b3e 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/QuasiCompact.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/QuasiCompact.lean @@ -38,26 +38,26 @@ of quasi-compact open sets are quasi-compact. @[mk_iff] class QuasiCompact (f : X ⟶ Y) : Prop where /-- Preimage of compact open set under a quasi-compact morphism between schemes is compact. -/ - isCompact_preimage : ∀ U : Set Y, IsOpen U → IsCompact U → IsCompact (f.1.base ⁻¹' U) + isCompact_preimage : ∀ U : Set Y, IsOpen U → IsCompact U → IsCompact (f.base ⁻¹' U) -theorem quasiCompact_iff_spectral : QuasiCompact f ↔ IsSpectralMap f.1.base := +theorem quasiCompact_iff_spectral : QuasiCompact f ↔ IsSpectralMap f.base := ⟨fun ⟨h⟩ => ⟨by fun_prop, h⟩, fun h => ⟨h.2⟩⟩ instance (priority := 900) quasiCompact_of_isIso {X Y : Scheme} (f : X ⟶ Y) [IsIso f] : QuasiCompact f := by constructor intro U _ hU' - convert hU'.image (inv f.1.base).continuous_toFun using 1 + convert hU'.image (inv f.base).continuous_toFun using 1 rw [Set.image_eq_preimage_of_inverse] · delta Function.LeftInverse - exact IsIso.inv_hom_id_apply f.1.base - · exact IsIso.hom_inv_id_apply f.1.base + exact IsIso.inv_hom_id_apply f.base + · exact IsIso.hom_inv_id_apply f.base instance quasiCompact_comp {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) [QuasiCompact f] [QuasiCompact g] : QuasiCompact (f ≫ g) := by constructor intro U hU hU' - rw [Scheme.comp_val_base, TopCat.coe_comp, Set.preimage_comp] + rw [Scheme.comp_base, TopCat.coe_comp, Set.preimage_comp] apply QuasiCompact.isCompact_preimage · exact Continuous.isOpen_preimage (by fun_prop) _ hU apply QuasiCompact.isCompact_preimage <;> assumption @@ -126,9 +126,9 @@ instance : HasAffineProperty @QuasiCompact (fun X _ _ _ ↦ CompactSpace X) wher isLocal_affineProperty := by constructor · apply AffineTargetMorphismProperty.respectsIso_mk <;> rintro X Y Z e _ _ H - exacts [@Homeomorph.compactSpace _ _ _ _ H (TopCat.homeoOfIso (asIso e.inv.1.base)), H] + exacts [@Homeomorph.compactSpace _ _ _ _ H (TopCat.homeoOfIso (asIso e.inv.base)), H] · introv _ H - change CompactSpace ((Opens.map f.val.base).obj (Y.basicOpen r)) + change CompactSpace ((Opens.map f.base).obj (Y.basicOpen r)) rw [Scheme.preimage_basicOpen f r] erw [← isCompact_iff_compactSpace] rw [← isCompact_univ_iff] at H @@ -137,7 +137,7 @@ instance : HasAffineProperty @QuasiCompact (fun X _ _ _ ↦ CompactSpace X) wher · rintro X Y H f S hS hS' rw [← IsAffineOpen.basicOpen_union_eq_self_iff] at hS · rw [← isCompact_univ_iff] - change IsCompact ((Opens.map f.val.base).obj ⊤).1 + change IsCompact ((Opens.map f.base).obj ⊤).1 rw [← hS] dsimp [Opens.map] simp only [Opens.iSup_mk, Opens.coe_mk, Set.preimage_iUnion] diff --git a/Mathlib/AlgebraicGeometry/Morphisms/QuasiSeparated.lean b/Mathlib/AlgebraicGeometry/Morphisms/QuasiSeparated.lean index e9ffc71cd889f..be7e03dfc20b1 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/QuasiSeparated.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/QuasiSeparated.lean @@ -351,7 +351,7 @@ theorem is_localization_basicOpen_of_qcqs {X : Scheme} {U : X.Opens} (hU : IsCom This is known as the **Qcqs lemma** in [R. Vakil, *The rising sea*][RisingSea]. -/ theorem isIso_ΓSpec_adjunction_unit_app_basicOpen {X : Scheme} [CompactSpace X] [QuasiSeparatedSpace X] (f : X.presheaf.obj (op ⊤)) : - IsIso ((ΓSpec.adjunction.unit.app X).val.c.app (op (PrimeSpectrum.basicOpen f))) := by + IsIso ((ΓSpec.adjunction.unit.app X).c.app (op (PrimeSpectrum.basicOpen f))) := by refine @IsIso.of_isIso_comp_right _ _ _ _ _ _ (X.presheaf.map (eqToHom (Scheme.toSpecΓ_preimage_basicOpen _ _).symm).op) _ ?_ rw [ConcreteCategory.isIso_iff_bijective, CommRingCat.forget_map] diff --git a/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean b/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean index 93a2ff38a6b0a..8e9613ac56f3a 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean @@ -438,7 +438,7 @@ lemma iff_exists_appLE_locally (hQi : RespectsIso Q) [HasRingHomProperty P (Loca Q (f.appLE U V e) := by refine ⟨fun hf x ↦ ?_, fun hf ↦ (IsLocalAtSource.iff_exists_resLE (P := P)).mpr <| fun x ↦ ?_⟩ · obtain ⟨U, hU, hfx, _⟩ := Opens.isBasis_iff_nbhd.mp (isBasis_affine_open Y) - (Opens.mem_top <| f.val.base x) + (Opens.mem_top <| f.base x) obtain ⟨V, hV, hx, e⟩ := Opens.isBasis_iff_nbhd.mp (isBasis_affine_open X) (show x ∈ f ⁻¹ᵁ U from hfx) simp_rw [HasRingHomProperty.iff_appLE (P := P), locally_iff_isLocalization hQi] at hf diff --git a/Mathlib/AlgebraicGeometry/Morphisms/UnderlyingMap.lean b/Mathlib/AlgebraicGeometry/Morphisms/UnderlyingMap.lean index b2c7a9ad87681..ff349f67ab6a8 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/UnderlyingMap.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/UnderlyingMap.lean @@ -51,12 +51,12 @@ variable {X Y Z : Scheme.{u}} (f : X ⟶ Y) (g : Y ⟶ Z) /-- A morphism of schemes is surjective if the underlying map is. -/ @[mk_iff] class Surjective : Prop where - surj : Function.Surjective f.1.base + surj : Function.Surjective f.base lemma surjective_eq_topologically : @Surjective = topologically Function.Surjective := by ext; exact surjective_iff _ -lemma Scheme.Hom.surjective (f : X.Hom Y) [Surjective f] : Function.Surjective f.1.base := +lemma Scheme.Hom.surjective (f : X.Hom Y) [Surjective f] : Function.Surjective f.base := Surjective.surj instance (priority := 100) [IsIso f] : Surjective f := ⟨f.homeomorph.surjective⟩ @@ -64,7 +64,7 @@ instance (priority := 100) [IsIso f] : Surjective f := ⟨f.homeomorph.surjectiv instance [Surjective f] [Surjective g] : Surjective (f ≫ g) := ⟨g.surjective.comp f.surjective⟩ lemma Surjective.of_comp [Surjective (f ≫ g)] : Surjective g where - surj := Function.Surjective.of_comp (g := f.1.base) (f ≫ g).surjective + surj := Function.Surjective.of_comp (g := f.base) (f ≫ g).surjective lemma Surjective.comp_iff [Surjective f] : Surjective (f ≫ g) ↔ Surjective g := ⟨fun _ ↦ of_comp f g, fun _ ↦ inferInstance⟩ diff --git a/Mathlib/AlgebraicGeometry/Morphisms/UniversallyClosed.lean b/Mathlib/AlgebraicGeometry/Morphisms/UniversallyClosed.lean index 3e0a0c0b91874..6a5489d992476 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/UniversallyClosed.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/UniversallyClosed.lean @@ -48,7 +48,7 @@ theorem universallyClosed_stableUnderBaseChange : StableUnderBaseChange @Univers instance isClosedMap_isStableUnderComposition : IsStableUnderComposition (topologically @IsClosedMap) where - comp_mem f g hf hg := IsClosedMap.comp (f := f.1.base) (g := g.1.base) hg hf + comp_mem f g hf hg := IsClosedMap.comp (f := f.base) (g := g.base) hg hf instance universallyClosed_isStableUnderComposition : IsStableUnderComposition @UniversallyClosed := by @@ -71,7 +71,7 @@ instance universallyClosed_isLocalAtTarget : IsLocalAtTarget @UniversallyClosed rw [universallyClosed_eq] apply universally_isLocalAtTarget intro X Y f ι U hU H - simp_rw [topologically, morphismRestrict_val_base] at H + simp_rw [topologically, morphismRestrict_base] at H exact (isClosedMap_iff_isClosedMap_of_iSup_eq_top hU).mpr H end AlgebraicGeometry diff --git a/Mathlib/AlgebraicGeometry/Noetherian.lean b/Mathlib/AlgebraicGeometry/Noetherian.lean index 5b47926dddcf1..45c741d358319 100644 --- a/Mathlib/AlgebraicGeometry/Noetherian.lean +++ b/Mathlib/AlgebraicGeometry/Noetherian.lean @@ -274,7 +274,7 @@ instance (priority := 100) IsNoetherian.noetherianSpace [IsNoetherian X] : apply TopologicalSpace.noetherian_univ_iff.mp let 𝒰 := X.affineCover.finiteSubcover rw [← 𝒰.iUnion_range] - suffices ∀ i : 𝒰.J, NoetherianSpace (Set.range <| (𝒰.map i).val.base) by + suffices ∀ i : 𝒰.J, NoetherianSpace (Set.range <| (𝒰.map i).base) by apply NoetherianSpace.iUnion intro i have : IsAffine (𝒰.obj i) := by diff --git a/Mathlib/AlgebraicGeometry/OpenImmersion.lean b/Mathlib/AlgebraicGeometry/OpenImmersion.lean index 860d3cba578d7..75394dff9384f 100644 --- a/Mathlib/AlgebraicGeometry/OpenImmersion.lean +++ b/Mathlib/AlgebraicGeometry/OpenImmersion.lean @@ -31,11 +31,11 @@ variable {C : Type u} [Category.{v} C] of LocallyRingedSpaces -/ abbrev IsOpenImmersion {X Y : Scheme.{u}} (f : X ⟶ Y) : Prop := - LocallyRingedSpace.IsOpenImmersion f + LocallyRingedSpace.IsOpenImmersion f.toLRSHom instance IsOpenImmersion.comp {X Y Z : Scheme.{u}} (f : X ⟶ Y) (g : Y ⟶ Z) [IsOpenImmersion f] [IsOpenImmersion g] : IsOpenImmersion (f ≫ g) := -LocallyRingedSpace.IsOpenImmersion.comp f g +LocallyRingedSpace.IsOpenImmersion.comp f.toLRSHom g.toLRSHom namespace LocallyRingedSpace.IsOpenImmersion @@ -45,7 +45,7 @@ protected def scheme (X : LocallyRingedSpace.{u}) (h : ∀ x : X, ∃ (R : CommRingCat) (f : Spec.toLocallyRingedSpace.obj (op R) ⟶ X), - (x ∈ Set.range f.1.base : _) ∧ LocallyRingedSpace.IsOpenImmersion f) : + (x ∈ Set.range f.base : _) ∧ LocallyRingedSpace.IsOpenImmersion f) : Scheme where toLocallyRingedSpace := X local_affine := by @@ -61,7 +61,7 @@ protected def scheme (X : LocallyRingedSpace.{u}) end LocallyRingedSpace.IsOpenImmersion theorem IsOpenImmersion.isOpen_range {X Y : Scheme.{u}} (f : X ⟶ Y) [H : IsOpenImmersion f] : - IsOpen (Set.range f.1.base) := + IsOpen (Set.range f.base) := H.base_open.isOpen_range @[deprecated (since := "2024-03-17")] @@ -71,7 +71,7 @@ namespace Scheme.Hom variable {X Y : Scheme.{u}} (f : Scheme.Hom X Y) [H : IsOpenImmersion f] -theorem openEmbedding : OpenEmbedding f.1.base := +theorem openEmbedding : OpenEmbedding f.base := H.base_open /-- The image of an open immersion as an open set. -/ @@ -81,14 +81,14 @@ def opensRange : Y.Opens := /-- The functor `opens X ⥤ opens Y` associated with an open immersion `f : X ⟶ Y`. -/ abbrev opensFunctor : X.Opens ⥤ Y.Opens := - LocallyRingedSpace.IsOpenImmersion.opensFunctor f + LocallyRingedSpace.IsOpenImmersion.opensFunctor f.toLRSHom /-- `f ''ᵁ U` is notation for the image (as an open set) of `U` under an open immersion `f`. -/ scoped[AlgebraicGeometry] notation3:90 f:91 " ''ᵁ " U:90 => (Scheme.Hom.opensFunctor f).obj U lemma image_le_image_of_le {U V : X.Opens} (e : U ≤ V) : f ''ᵁ U ≤ f ''ᵁ V := by rintro a ⟨u, hu, rfl⟩ - exact Set.mem_image_of_mem (⇑f.val.base) (e hu) + exact Set.mem_image_of_mem (⇑f.base) (e hu) @[simp] lemma opensFunctor_map_homOfLE {U V : X.Opens} (e : U ≤ V) : @@ -131,7 +131,7 @@ lemma image_iSup₂ {ι : Sort*} {κ : ι → Sort*} (s : (i : ι) → κ i → /-- The isomorphism `Γ(Y, f(U)) ≅ Γ(X, U)` induced by an open immersion `f : X ⟶ Y`. -/ def appIso (U) : Γ(Y, f ''ᵁ U) ≅ Γ(X, U) := - (asIso <| LocallyRingedSpace.IsOpenImmersion.invApp f U).symm + (asIso <| LocallyRingedSpace.IsOpenImmersion.invApp f.toLRSHom U).symm @[reassoc (attr := simp)] theorem appIso_inv_naturality {U V : X.Opens} (i : op U ⟶ op V) : @@ -142,7 +142,7 @@ theorem appIso_inv_naturality {U V : X.Opens} (i : op U ⟶ op V) : theorem appIso_hom (U) : (f.appIso U).hom = f.app (f ''ᵁ U) ≫ X.presheaf.map (eqToHom (preimage_image_eq f U).symm).op := - (PresheafedSpace.IsOpenImmersion.inv_invApp f.1 U).trans (by rw [eqToHom_op]) + (PresheafedSpace.IsOpenImmersion.inv_invApp f.toPshHom U).trans (by rw [eqToHom_op]) theorem appIso_hom' (U) : (f.appIso U).hom = f.appLE (f ''ᵁ U) U (preimage_image_eq f U).ge := @@ -151,7 +151,7 @@ theorem appIso_hom' (U) : @[reassoc (attr := simp)] theorem app_appIso_inv (U) : f.app U ≫ (f.appIso (f ⁻¹ᵁ U)).inv = - Y.presheaf.map (homOfLE (Set.image_preimage_subset f.1.base U.1)).op := + Y.presheaf.map (homOfLE (Set.image_preimage_subset f.base U.1)).op := PresheafedSpace.IsOpenImmersion.app_invApp _ _ /-- A variant of `app_invApp` that gives an `eqToHom` instead of `homOfLE`. -/ @@ -227,19 +227,18 @@ lemma _root_.AlgebraicGeometry.IsOpenImmersion.of_isLocalization {R S} [CommRing theorem exists_affine_mem_range_and_range_subset {X : Scheme.{u}} {x : X} {U : X.Opens} (hxU : x ∈ U) : ∃ (R : CommRingCat) (f : Spec R ⟶ X), - IsOpenImmersion f ∧ x ∈ Set.range f.1.base ∧ Set.range f.1.base ⊆ U := by + IsOpenImmersion f ∧ x ∈ Set.range f.base ∧ Set.range f.base ⊆ U := by obtain ⟨⟨V, hxV⟩, R, ⟨e⟩⟩ := X.2 x - have : e.hom.1.base ⟨x, hxV⟩ ∈ (Opens.map (e.inv.1.base ≫ V.inclusion')).obj U := - show ((e.hom ≫ e.inv).1.base ⟨x, hxV⟩).1 ∈ U from e.hom_inv_id ▸ hxU + have : e.hom.base ⟨x, hxV⟩ ∈ (Opens.map (e.inv.base ≫ V.inclusion')).obj U := + show ((e.hom ≫ e.inv).base ⟨x, hxV⟩).1 ∈ U from e.hom_inv_id ▸ hxU obtain ⟨_, ⟨_, ⟨r : R, rfl⟩, rfl⟩, hr, hr'⟩ := PrimeSpectrum.isBasis_basic_opens.exists_subset_of_mem_open this (Opens.is_open' _) let f : Spec (CommRingCat.of (Localization.Away r)) ⟶ X := - Spec.map (CommRingCat.ofHom (algebraMap R (Localization.Away r))) ≫ (e.inv ≫ X.ofRestrict _ : _) + Spec.map (CommRingCat.ofHom (algebraMap R (Localization.Away r))) ≫ ⟨e.inv ≫ X.ofRestrict _⟩ refine ⟨.of (Localization.Away r), f, inferInstance, ?_⟩ - rw [Scheme.comp_val_base, LocallyRingedSpace.comp_val, SheafedSpace.comp_base, TopCat.coe_comp, - Set.range_comp] + rw [Scheme.comp_base, TopCat.coe_comp, Set.range_comp] erw [PrimeSpectrum.localization_away_comap_range (Localization.Away r) r] - exact ⟨⟨_, hr, congr(($(e.hom_inv_id).1.base ⟨x, hxV⟩).1)⟩, Set.image_subset_iff.mpr hr'⟩ + exact ⟨⟨_, hr, congr(($(e.hom_inv_id).base ⟨x, hxV⟩).1)⟩, Set.image_subset_iff.mpr hr'⟩ end Scheme @@ -269,10 +268,10 @@ theorem toScheme_toLocallyRingedSpace : upgrade it into a morphism of Schemes. -/ def toSchemeHom : toScheme Y f ⟶ Y := - toLocallyRingedSpaceHom _ f + ⟨toLocallyRingedSpaceHom _ f⟩ @[simp] -theorem toSchemeHom_val : (toSchemeHom Y f).val = f := +theorem toSchemeHom_toPshHom : (toSchemeHom Y f).toPshHom = f := rfl instance toSchemeHom_isOpenImmersion : AlgebraicGeometry.IsOpenImmersion (toSchemeHom Y f) := @@ -283,9 +282,9 @@ theorem scheme_eq_of_locallyRingedSpace_eq {X Y : Scheme.{u}} cases X; cases Y; congr theorem scheme_toScheme {X Y : Scheme.{u}} (f : X ⟶ Y) [AlgebraicGeometry.IsOpenImmersion f] : - toScheme Y f.1 = X := by + toScheme Y f.toPshHom = X := by apply scheme_eq_of_locallyRingedSpace_eq - exact locallyRingedSpace_toLocallyRingedSpace f + exact locallyRingedSpace_toLocallyRingedSpace f.toLRSHom end ToScheme @@ -305,14 +304,14 @@ lemma Scheme.restrict_toPresheafedSpace : (X.restrict h).toPresheafedSpace = X.toPresheafedSpace.restrict h := rfl /-- The canonical map from the restriction to the subspace. -/ -@[simps! val_base, simps! (config := .lemmasOnly) val_c_app] +@[simps! toLRSHom_base, simps! (config := .lemmasOnly) toLRSHom_c_app] def Scheme.ofRestrict : X.restrict h ⟶ X := - X.toLocallyRingedSpace.ofRestrict h + ⟨X.toLocallyRingedSpace.ofRestrict h⟩ @[simp] lemma Scheme.ofRestrict_app (V) : (X.ofRestrict h).app V = X.presheaf.map (h.isOpenMap.adjunction.counit.app V).op := - Scheme.ofRestrict_val_c_app X h (op V) + Scheme.ofRestrict_toLRSHom_c_app X h (op V) instance IsOpenImmersion.ofRestrict : IsOpenImmersion (X.ofRestrict h) := show PresheafedSpace.IsOpenImmersion (X.toPresheafedSpace.ofRestrict h) by infer_instance @@ -344,45 +343,44 @@ variable {X Y Z : Scheme.{u}} (f : X ⟶ Z) (g : Y ⟶ Z) variable [H : IsOpenImmersion f] instance (priority := 100) of_isIso [IsIso g] : IsOpenImmersion g := - @LocallyRingedSpace.IsOpenImmersion.of_isIso _ _ _ - (show IsIso ((inducedFunctor _).map g) by infer_instance) + LocallyRingedSpace.IsOpenImmersion.of_isIso _ -theorem to_iso {X Y : Scheme.{u}} (f : X ⟶ Y) [h : IsOpenImmersion f] [Epi f.1.base] : IsIso f := +theorem to_iso {X Y : Scheme.{u}} (f : X ⟶ Y) [h : IsOpenImmersion f] [Epi f.base] : IsIso f := @isIso_of_reflects_iso _ _ _ _ _ _ f (Scheme.forgetToLocallyRingedSpace ⋙ LocallyRingedSpace.forgetToSheafedSpace ⋙ SheafedSpace.forgetToPresheafedSpace) - (@PresheafedSpace.IsOpenImmersion.to_iso _ _ _ _ f.1 h _) _ + (@PresheafedSpace.IsOpenImmersion.to_iso _ _ _ _ f.toPshHom h _) _ -theorem of_stalk_iso {X Y : Scheme.{u}} (f : X ⟶ Y) (hf : OpenEmbedding f.1.base) +theorem of_stalk_iso {X Y : Scheme.{u}} (f : X ⟶ Y) (hf : OpenEmbedding f.base) [∀ x, IsIso (f.stalkMap x)] : IsOpenImmersion f := - haveI (x : X) : IsIso (f.val.stalkMap x) := inferInstanceAs <| IsIso (f.stalkMap x) - SheafedSpace.IsOpenImmersion.of_stalk_iso f.1 hf + haveI (x : X) : IsIso (f.toShHom.stalkMap x) := inferInstanceAs <| IsIso (f.stalkMap x) + SheafedSpace.IsOpenImmersion.of_stalk_iso f.toShHom hf instance stalk_iso {X Y : Scheme.{u}} (f : X ⟶ Y) [IsOpenImmersion f] (x : X) : IsIso (f.stalkMap x) := - inferInstanceAs <| IsIso (f.val.stalkMap x) + inferInstanceAs <| IsIso (f.toLRSHom.stalkMap x) lemma of_comp {X Y Z : Scheme.{u}} (f : X ⟶ Y) (g : Y ⟶ Z) [IsOpenImmersion g] [IsOpenImmersion (f ≫ g)] : IsOpenImmersion f := haveI (x : X) : IsIso (f.stalkMap x) := - haveI : IsIso (g.stalkMap (f.val.base x) ≫ f.stalkMap x) := by + haveI : IsIso (g.stalkMap (f.base x) ≫ f.stalkMap x) := by rw [← Scheme.stalkMap_comp] infer_instance - IsIso.of_isIso_comp_left (f := g.stalkMap (f.val.base x)) _ + IsIso.of_isIso_comp_left (f := g.stalkMap (f.base x)) _ IsOpenImmersion.of_stalk_iso _ <| OpenEmbedding.of_comp _ (Scheme.Hom.openEmbedding g) (Scheme.Hom.openEmbedding (f ≫ g)) theorem iff_stalk_iso {X Y : Scheme.{u}} (f : X ⟶ Y) : - IsOpenImmersion f ↔ OpenEmbedding f.1.base ∧ ∀ x, IsIso (f.stalkMap x) := - ⟨fun H => ⟨H.1, fun x ↦ inferInstanceAs <| IsIso (f.val.stalkMap x)⟩, + IsOpenImmersion f ↔ OpenEmbedding f.base ∧ ∀ x, IsIso (f.stalkMap x) := + ⟨fun H => ⟨H.1, fun x ↦ inferInstanceAs <| IsIso (f.toPshHom.stalkMap x)⟩, fun ⟨h₁, h₂⟩ => @IsOpenImmersion.of_stalk_iso _ _ f h₁ h₂⟩ theorem _root_.AlgebraicGeometry.isIso_iff_isOpenImmersion {X Y : Scheme.{u}} (f : X ⟶ Y) : - IsIso f ↔ IsOpenImmersion f ∧ Epi f.1.base := + IsIso f ↔ IsOpenImmersion f ∧ Epi f.base := ⟨fun _ => ⟨inferInstance, inferInstance⟩, fun ⟨h₁, h₂⟩ => @IsOpenImmersion.to_iso _ _ f h₁ h₂⟩ theorem _root_.AlgebraicGeometry.isIso_iff_stalk_iso {X Y : Scheme.{u}} (f : X ⟶ Y) : - IsIso f ↔ IsIso f.1.base ∧ ∀ x, IsIso (f.stalkMap x) := by + IsIso f ↔ IsIso f.base ∧ ∀ x, IsIso (f.stalkMap x) := by rw [isIso_iff_isOpenImmersion, IsOpenImmersion.iff_stalk_iso, and_comm, ← and_assoc] refine and_congr ⟨?_, ?_⟩ Iff.rfl · rintro ⟨h₁, h₂⟩ @@ -393,16 +391,18 @@ theorem _root_.AlgebraicGeometry.isIso_iff_stalk_iso {X Y : Scheme.{u}} (f : X (Equiv.ofBijective _ ⟨h₂.inj, (TopCat.epi_iff_surjective _).mp h₁⟩) h₂.continuous h₂.isOpenMap)).hom infer_instance - · intro H; exact ⟨inferInstance, (TopCat.homeoOfIso (asIso f.1.base)).openEmbedding⟩ + · intro H; exact ⟨inferInstance, (TopCat.homeoOfIso (asIso f.base)).openEmbedding⟩ /-- An open immersion induces an isomorphism from the domain onto the image -/ -def isoRestrict : X ≅ (Z.restrict H.base_open : _) where - __ := (LocallyRingedSpace.IsOpenImmersion.isoRestrict f) +def isoRestrict : X ≅ (Z.restrict H.base_open : _) := + Scheme.fullyFaithfulForgetToLocallyRingedSpace.preimageIso + (LocallyRingedSpace.IsOpenImmersion.isoRestrict f.toLRSHom) local notation "forget" => Scheme.forgetToLocallyRingedSpace instance mono : Mono f := - (inducedFunctor _).mono_of_mono_map (show @Mono LocallyRingedSpace _ _ _ f by infer_instance) + Scheme.forgetToLocallyRingedSpace.mono_of_mono_map + (show Mono f.toLRSHom by infer_instance) instance forget_map_isOpenImmersion : LocallyRingedSpace.IsOpenImmersion ((forget).map f) := ⟨H.base_open, H.c_iso⟩ @@ -430,12 +430,12 @@ instance hasLimit_cospan_forget_of_right' : instance forgetCreatesPullbackOfLeft : CreatesLimit (cospan f g) forget := createsLimitOfFullyFaithfulOfIso - (PresheafedSpace.IsOpenImmersion.toScheme Y (@pullback.snd LocallyRingedSpace _ _ _ _ f g _).1) + (PresheafedSpace.IsOpenImmersion.toScheme Y (pullback.snd f.toLRSHom g.toLRSHom).toShHom) (eqToIso (by simp) ≪≫ HasLimit.isoOfNatIso (diagramIsoCospan _).symm) instance forgetCreatesPullbackOfRight : CreatesLimit (cospan g f) forget := createsLimitOfFullyFaithfulOfIso - (PresheafedSpace.IsOpenImmersion.toScheme Y (@pullback.fst LocallyRingedSpace _ _ _ _ g f _).1) + (PresheafedSpace.IsOpenImmersion.toScheme Y (pullback.fst g.toLRSHom f.toLRSHom).1) (eqToIso (by simp) ≪≫ HasLimit.isoOfNatIso (diagramIsoCospan _).symm) instance forgetPreservesOfLeft : PreservesLimit (cospan f g) forget := @@ -453,8 +453,8 @@ instance hasPullback_of_right : HasPullback g f := instance pullback_snd_of_left : IsOpenImmersion (pullback.snd f g) := by have := PreservesPullback.iso_hom_snd forget f g dsimp only [Scheme.forgetToLocallyRingedSpace, inducedFunctor_map] at this - rw [← this] change LocallyRingedSpace.IsOpenImmersion _ + rw [← this] infer_instance instance pullback_fst_of_right : IsOpenImmersion (pullback.fst g f) := by @@ -484,10 +484,10 @@ instance forgetToTopPreservesOfRight : PreservesLimit (cospan g f) Scheme.forget preservesPullbackSymmetry _ _ _ theorem range_pullback_snd_of_left : - Set.range (pullback.snd f g).1.base = (g ⁻¹ᵁ f.opensRange).1 := by - rw [← show _ = (pullback.snd f g).1.base from + Set.range (pullback.snd f g).base = (g ⁻¹ᵁ f.opensRange).1 := by + rw [← show _ = (pullback.snd f g).base from PreservesPullback.iso_hom_snd Scheme.forgetToTop f g, TopCat.coe_comp, Set.range_comp, - Set.range_iff_surjective.mpr, ← @Set.preimage_univ _ _ (pullback.fst f.1.base g.1.base)] + Set.range_iff_surjective.mpr, ← @Set.preimage_univ _ _ (pullback.fst f.base g.base)] -- Porting note (#10691): was `rw` · erw [TopCat.pullback_snd_image_fst_preimage] rw [Set.image_univ] @@ -500,11 +500,11 @@ theorem opensRange_pullback_snd_of_left : Opens.ext (range_pullback_snd_of_left f g) theorem range_pullback_fst_of_right : - Set.range (pullback.fst g f).1.base = - ((Opens.map g.1.base).obj ⟨Set.range f.1.base, H.base_open.isOpen_range⟩).1 := by - rw [← show _ = (pullback.fst g f).1.base from + Set.range (pullback.fst g f).base = + ((Opens.map g.base).obj ⟨Set.range f.base, H.base_open.isOpen_range⟩).1 := by + rw [← show _ = (pullback.fst g f).base from PreservesPullback.iso_hom_fst Scheme.forgetToTop g f, TopCat.coe_comp, Set.range_comp, - Set.range_iff_surjective.mpr, ← @Set.preimage_univ _ _ (pullback.snd g.1.base f.1.base)] + Set.range_iff_surjective.mpr, ← @Set.preimage_univ _ _ (pullback.snd g.base f.base)] -- Porting note (#10691): was `rw` · erw [TopCat.pullback_fst_image_snd_preimage] rw [Set.image_univ] @@ -517,16 +517,16 @@ theorem opensRange_pullback_fst_of_right : Opens.ext (range_pullback_fst_of_right f g) theorem range_pullback_to_base_of_left : - Set.range (pullback.fst f g ≫ f).1.base = - Set.range f.1.base ∩ Set.range g.1.base := by - rw [pullback.condition, Scheme.comp_val_base, TopCat.coe_comp, Set.range_comp, + Set.range (pullback.fst f g ≫ f).base = + Set.range f.base ∩ Set.range g.base := by + rw [pullback.condition, Scheme.comp_base, TopCat.coe_comp, Set.range_comp, range_pullback_snd_of_left, Opens.carrier_eq_coe, Opens.map_obj, Opens.coe_mk, Set.image_preimage_eq_inter_range, Opens.carrier_eq_coe, Scheme.Hom.opensRange_coe] theorem range_pullback_to_base_of_right : - Set.range (pullback.fst g f ≫ g).1.base = - Set.range g.1.base ∩ Set.range f.1.base := by - rw [Scheme.comp_val_base, TopCat.coe_comp, Set.range_comp, range_pullback_fst_of_right, + Set.range (pullback.fst g f ≫ g).base = + Set.range g.base ∩ Set.range f.base := by + rw [Scheme.comp_base, TopCat.coe_comp, Set.range_comp, range_pullback_fst_of_right, Opens.map_obj, Opens.carrier_eq_coe, Opens.coe_mk, Set.image_preimage_eq_inter_range, Set.inter_comm] @@ -535,19 +535,20 @@ For an open immersion `f : X ⟶ Z`, given any morphism of schemes `g : Y ⟶ Z` image is contained in the image of `f`, we can lift this morphism to a unique `Y ⟶ X` that commutes with these maps. -/ -def lift (H' : Set.range g.1.base ⊆ Set.range f.1.base) : Y ⟶ X := - LocallyRingedSpace.IsOpenImmersion.lift f g H' +def lift (H' : Set.range g.base ⊆ Set.range f.base) : Y ⟶ X := + ⟨LocallyRingedSpace.IsOpenImmersion.lift f.toLRSHom g.toLRSHom H'⟩ @[simp, reassoc] -theorem lift_fac (H' : Set.range g.1.base ⊆ Set.range f.1.base) : lift f g H' ≫ f = g := - LocallyRingedSpace.IsOpenImmersion.lift_fac f g H' +theorem lift_fac (H' : Set.range g.base ⊆ Set.range f.base) : lift f g H' ≫ f = g := + Scheme.Hom.ext' <| LocallyRingedSpace.IsOpenImmersion.lift_fac f.toLRSHom g.toLRSHom H' -theorem lift_uniq (H' : Set.range g.1.base ⊆ Set.range f.1.base) (l : Y ⟶ X) (hl : l ≫ f = g) : +theorem lift_uniq (H' : Set.range g.base ⊆ Set.range f.base) (l : Y ⟶ X) (hl : l ≫ f = g) : l = lift f g H' := - LocallyRingedSpace.IsOpenImmersion.lift_uniq f g H' l hl + Scheme.Hom.ext' <| LocallyRingedSpace.IsOpenImmersion.lift_uniq + f.toLRSHom g.toLRSHom H' l.toLRSHom congr(($hl).toLRSHom) /-- Two open immersions with equal range are isomorphic. -/ -def isoOfRangeEq [IsOpenImmersion g] (e : Set.range f.1.base = Set.range g.1.base) : X ≅ Y where +def isoOfRangeEq [IsOpenImmersion g] (e : Set.range f.base = Set.range g.base) : X ≅ Y where hom := lift g f (le_of_eq e) inv := lift f g (le_of_eq e.symm) hom_inv_id := by rw [← cancel_mono f]; simp @@ -555,13 +556,13 @@ def isoOfRangeEq [IsOpenImmersion g] (e : Set.range f.1.base = Set.range g.1.bas @[simp, reassoc] lemma isoOfRangeEq_hom_fac {X Y Z : Scheme.{u}} (f : X ⟶ Z) (g : Y ⟶ Z) - [IsOpenImmersion f] [IsOpenImmersion g] (e : Set.range f.1.base = Set.range g.1.base) : + [IsOpenImmersion f] [IsOpenImmersion g] (e : Set.range f.base = Set.range g.base) : (isoOfRangeEq f g e).hom ≫ g = f := lift_fac _ _ (le_of_eq e) @[simp, reassoc] lemma isoOfRangeEq_inv_fac {X Y Z : Scheme.{u}} (f : X ⟶ Z) (g : Y ⟶ Z) - [IsOpenImmersion f] [IsOpenImmersion g] (e : Set.range f.1.base = Set.range g.1.base) : + [IsOpenImmersion f] [IsOpenImmersion g] (e : Set.range f.base = Set.range g.base) : (isoOfRangeEq f g e).inv ≫ f = g := lift_fac _ _ (le_of_eq e.symm) @@ -569,7 +570,7 @@ theorem app_eq_invApp_app_of_comp_eq_aux {X Y U : Scheme.{u}} (f : Y ⟶ U) (g : (H : fg = f ≫ g) [h : IsOpenImmersion g] (V : U.Opens) : f ⁻¹ᵁ V = fg ⁻¹ᵁ (g ''ᵁ V) := by subst H - rw [Scheme.comp_val_base, Opens.map_comp_obj] + rw [Scheme.comp_base, Opens.map_comp_obj] congr 1 ext1 exact (Set.preimage_image_eq _ h.base_open.inj).symm @@ -625,7 +626,7 @@ def ΓIsoTop {X Y : Scheme.{u}} (f : X ⟶ Y) [IsOpenImmersion f] : (f.appIso ⊤).symm ≪≫ Y.presheaf.mapIso (eqToIso f.image_top_eq_opensRange.symm).op instance {Z : Scheme.{u}} (f : X ⟶ Z) (g : Y ⟶ Z) [IsOpenImmersion f] - (H' : Set.range g.val.base ⊆ Set.range f.val.base) [IsOpenImmersion g] : + (H' : Set.range g.base ⊆ Set.range f.base) [IsOpenImmersion g] : IsOpenImmersion (IsOpenImmersion.lift f g H') := haveI : IsOpenImmersion (IsOpenImmersion.lift f g H' ≫ f) := by simpa IsOpenImmersion.of_comp _ f diff --git a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean index 35a672056282e..106326628a9e4 100644 --- a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean +++ b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean @@ -635,7 +635,7 @@ def toSpec (f) : (Proj| pbo f) ⟶ Spec (A⁰_ f) := open HomogeneousLocalization LocalRing lemma toSpec_base_apply_eq_comap {f} (x : Proj| pbo f) : - (toSpec 𝒜 f).1.base x = PrimeSpectrum.comap (mapId 𝒜 (Submonoid.powers_le.mpr x.2)) + (toSpec 𝒜 f).base x = PrimeSpectrum.comap (mapId 𝒜 (Submonoid.powers_le.mpr x.2)) (closedPoint (AtPrime 𝒜 x.1.asHomogeneousIdeal.toIdeal)) := by show PrimeSpectrum.comap (awayToΓ 𝒜 f ≫ (Proj| pbo f).presheaf.Γgerm x) (LocalRing.closedPoint ((Proj| pbo f).presheaf.stalk x)) = _ @@ -645,7 +645,7 @@ lemma toSpec_base_apply_eq_comap {f} (x : Proj| pbo f) : ((Proj| pbo f).presheaf.stalk x) _ _ _ (isLocalRingHom_of_isIso _))) lemma toSpec_base_apply_eq {f} (x : Proj| pbo f) : - (toSpec 𝒜 f).1.base x = ProjIsoSpecTopComponent.toSpec 𝒜 f x := + (toSpec 𝒜 f).base x = ProjIsoSpecTopComponent.toSpec 𝒜 f x := toSpec_base_apply_eq_comap 𝒜 x |>.trans <| PrimeSpectrum.ext <| Ideal.ext fun z => show ¬ IsUnit _ ↔ z ∈ ProjIsoSpecTopComponent.ToSpec.carrier _ by obtain ⟨z, rfl⟩ := z.mk_surjective @@ -656,19 +656,19 @@ lemma toSpec_base_apply_eq {f} (x : Proj| pbo f) : exact not_not lemma toSpec_base_isIso {f} {m} (f_deg : f ∈ 𝒜 m) (hm : 0 < m) : - IsIso (toSpec 𝒜 f).1.base := by + IsIso (toSpec 𝒜 f).base := by convert (projIsoSpecTopComponent f_deg hm).isIso_hom exact DFunLike.ext _ _ <| toSpec_base_apply_eq 𝒜 lemma mk_mem_toSpec_base_apply {f} (x : Proj| pbo f) (z : NumDenSameDeg 𝒜 (.powers f)) : - HomogeneousLocalization.mk z ∈ ((toSpec 𝒜 f).1.base x).asIdeal ↔ + HomogeneousLocalization.mk z ∈ ((toSpec 𝒜 f).base x).asIdeal ↔ z.num.1 ∈ x.1.asHomogeneousIdeal := (toSpec_base_apply_eq 𝒜 x).symm ▸ ProjIsoSpecTopComponent.ToSpec.mk_mem_carrier _ _ lemma toSpec_preimage_basicOpen {f} (t : NumDenSameDeg 𝒜 (.powers f)) : - (Opens.map (toSpec 𝒜 f).1.base).obj (sbo (.mk t)) = + (Opens.map (toSpec 𝒜 f).base).obj (sbo (.mk t)) = Opens.comap ⟨_, continuous_subtype_val⟩ (pbo t.num.1) := Opens.ext <| Opens.map_coe _ _ ▸ by convert (ProjIsoSpecTopComponent.ToSpec.preimage_basicOpen f t) @@ -676,7 +676,7 @@ lemma toSpec_preimage_basicOpen {f} @[reassoc] lemma toOpen_toSpec_val_c_app (f) (U) : - StructureSheaf.toOpen (A⁰_ f) U.unop ≫ (toSpec 𝒜 f).val.c.app U = + StructureSheaf.toOpen (A⁰_ f) U.unop ≫ (toSpec 𝒜 f).c.app U = awayToΓ 𝒜 f ≫ (Proj| pbo f).presheaf.map (homOfLE le_top).op := Eq.trans (by congr) <| ΓSpec.toOpen_comp_locallyRingedSpaceAdjunction_homEquiv_app _ U @@ -697,7 +697,7 @@ of `A⁰_f` at `φ(x)` where `φ : Proj|D(f) ⟶ Spec A⁰_f` is the morphism of constructed as above. -/ lemma isLocalization_atPrime (f) (x : pbo f) {m} (f_deg : f ∈ 𝒜 m) (hm : 0 < m) : - @IsLocalization (Away 𝒜 f) _ ((toSpec 𝒜 f).1.base x).asIdeal.primeCompl + @IsLocalization (Away 𝒜 f) _ ((toSpec 𝒜 f).base x).asIdeal.primeCompl (AtPrime 𝒜 x.1.asHomogeneousIdeal.toIdeal) _ (mapId 𝒜 (Submonoid.powers_le.mpr x.2)).toAlgebra := by letI : Algebra (Away 𝒜 f) (AtPrime 𝒜 x.1.asHomogeneousIdeal.toIdeal) := @@ -755,46 +755,46 @@ stalk of `Spec A⁰_ f` at `y` is isomorphic to `A⁰ₓ` where `y` is the point to `x`. -/ def specStalkEquiv (f) (x : pbo f) {m} (f_deg : f ∈ 𝒜 m) (hm : 0 < m) : - (Spec.structureSheaf (A⁰_ f)).presheaf.stalk ((toSpec 𝒜 f).1.base x) ≅ + (Spec.structureSheaf (A⁰_ f)).presheaf.stalk ((toSpec 𝒜 f).base x) ≅ CommRingCat.of (AtPrime 𝒜 x.1.asHomogeneousIdeal.toIdeal) := letI : Algebra (Away 𝒜 f) (AtPrime 𝒜 x.1.asHomogeneousIdeal.toIdeal) := (mapId 𝒜 (Submonoid.powers_le.mpr x.2)).toAlgebra haveI := isLocalization_atPrime 𝒜 f x f_deg hm (IsLocalization.algEquiv (R := A⁰_ f) - (M := ((toSpec 𝒜 f).1.base x).asIdeal.primeCompl) - (S := (Spec.structureSheaf (A⁰_ f)).presheaf.stalk ((toSpec 𝒜 f).1.base x)) + (M := ((toSpec 𝒜 f).base x).asIdeal.primeCompl) + (S := (Spec.structureSheaf (A⁰_ f)).presheaf.stalk ((toSpec 𝒜 f).base x)) (Q := AtPrime 𝒜 x.1.asHomogeneousIdeal.toIdeal)).toRingEquiv.toCommRingCatIso lemma toStalk_specStalkEquiv (f) (x : pbo f) {m} (f_deg : f ∈ 𝒜 m) (hm : 0 < m) : - StructureSheaf.toStalk (A⁰_ f) ((toSpec 𝒜 f).1.base x) ≫ (specStalkEquiv 𝒜 f x f_deg hm).hom = + StructureSheaf.toStalk (A⁰_ f) ((toSpec 𝒜 f).base x) ≫ (specStalkEquiv 𝒜 f x f_deg hm).hom = (mapId _ <| Submonoid.powers_le.mpr x.2 : (A⁰_ f) →+* AtPrime 𝒜 x.1.1.toIdeal) := letI : Algebra (Away 𝒜 f) (AtPrime 𝒜 x.1.asHomogeneousIdeal.toIdeal) := (mapId 𝒜 (Submonoid.powers_le.mpr x.2)).toAlgebra letI := isLocalization_atPrime 𝒜 f x f_deg hm (IsLocalization.algEquiv (R := A⁰_ f) - (M := ((toSpec 𝒜 f).1.base x).asIdeal.primeCompl) - (S := (Spec.structureSheaf (A⁰_ f)).presheaf.stalk ((toSpec 𝒜 f).1.base x)) + (M := ((toSpec 𝒜 f).base x).asIdeal.primeCompl) + (S := (Spec.structureSheaf (A⁰_ f)).presheaf.stalk ((toSpec 𝒜 f).base x)) (Q := AtPrime 𝒜 x.1.asHomogeneousIdeal.toIdeal)).toAlgHom.comp_algebraMap lemma stalkMap_toSpec (f) (x : pbo f) {m} (f_deg : f ∈ 𝒜 m) (hm : 0 < m) : (toSpec 𝒜 f).stalkMap x = (specStalkEquiv 𝒜 f x f_deg hm).hom ≫ (Proj.stalkIso' 𝒜 x.1).toCommRingCatIso.inv ≫ ((Proj.toLocallyRingedSpace 𝒜).restrictStalkIso (Opens.openEmbedding _) x).inv := - IsLocalization.ringHom_ext (R := A⁰_ f) ((toSpec 𝒜 f).1.base x).asIdeal.primeCompl - (S := (Spec.structureSheaf (A⁰_ f)).presheaf.stalk ((toSpec 𝒜 f).1.base x)) <| + IsLocalization.ringHom_ext (R := A⁰_ f) ((toSpec 𝒜 f).base x).asIdeal.primeCompl + (S := (Spec.structureSheaf (A⁰_ f)).presheaf.stalk ((toSpec 𝒜 f).base x)) <| (toStalk_stalkMap_toSpec _ _ _).trans <| by rw [awayToΓ_ΓToStalk, ← toStalk_specStalkEquiv, Category.assoc]; rfl lemma isIso_toSpec (f) {m} (f_deg : f ∈ 𝒜 m) (hm : 0 < m) : IsIso (toSpec 𝒜 f) := by - haveI : IsIso (toSpec 𝒜 f).1.base := toSpec_base_isIso 𝒜 f_deg hm + haveI : IsIso (toSpec 𝒜 f).base := toSpec_base_isIso 𝒜 f_deg hm haveI (x) : IsIso ((toSpec 𝒜 f).stalkMap x) := by rw [stalkMap_toSpec 𝒜 f x f_deg hm]; infer_instance haveI : LocallyRingedSpace.IsOpenImmersion (toSpec 𝒜 f) := LocallyRingedSpace.IsOpenImmersion.of_stalk_iso (toSpec 𝒜 f) - (TopCat.homeoOfIso (asIso <| (toSpec 𝒜 f).1.base)).openEmbedding + (TopCat.homeoOfIso (asIso <| (toSpec 𝒜 f).base)).openEmbedding exact LocallyRingedSpace.IsOpenImmersion.to_iso _ end ProjectiveSpectrum.Proj diff --git a/Mathlib/AlgebraicGeometry/Properties.lean b/Mathlib/AlgebraicGeometry/Properties.lean index 626ba8ab0aecf..a5bbb08d90c15 100644 --- a/Mathlib/AlgebraicGeometry/Properties.lean +++ b/Mathlib/AlgebraicGeometry/Properties.lean @@ -29,16 +29,14 @@ namespace AlgebraicGeometry variable (X : Scheme) -instance : T0Space X := by - refine T0Space.of_open_cover fun x => ?_ - obtain ⟨U, R, ⟨e⟩⟩ := X.local_affine x - let e' : U.1 ≃ₜ PrimeSpectrum R := - homeoOfIso ((LocallyRingedSpace.forgetToSheafedSpace ⋙ SheafedSpace.forget _).mapIso e) - exact ⟨U.1.1, U.2, U.1.2, e'.embedding.t0Space⟩ +instance : T0Space X := + T0Space.of_open_cover fun x => ⟨_, X.affineCover.covers x, + (X.affineCover.map x).opensRange.2, Embedding.t0Space (Y := PrimeSpectrum _) + (isAffineOpen_opensRange (X.affineCover.map x)).isoSpec.schemeIsoToHomeo.embedding⟩ instance : QuasiSober X := by apply (config := { allowSynthFailures := true }) - quasiSober_of_open_cover (Set.range fun x => Set.range <| (X.affineCover.map x).1.base) + quasiSober_of_open_cover (Set.range fun x => Set.range <| (X.affineCover.map x).base) · rintro ⟨_, i, rfl⟩; exact (X.affineCover.IsOpen i).base_open.isOpen_range · rintro ⟨_, i, rfl⟩ exact @OpenEmbedding.quasiSober _ _ _ _ _ (Homeomorph.ofEmbedding _ @@ -128,7 +126,7 @@ theorem reduce_to_affine_global (P : ∀ {X : Scheme} (_ : X.Opens), Prop) theorem reduce_to_affine_nbhd (P : ∀ (X : Scheme) (_ : X), Prop) (h₁ : ∀ R x, P (Spec R) x) - (h₂ : ∀ {X Y} (f : X ⟶ Y) [IsOpenImmersion f] (x : X), P X x → P Y (f.1.base x)) : + (h₂ : ∀ {X Y} (f : X ⟶ Y) [IsOpenImmersion f] (x : X), P X x → P Y (f.base x)) : ∀ (X : Scheme) (x : X), P X x := by intro X x obtain ⟨y, e⟩ := X.affineCover.covers x diff --git a/Mathlib/AlgebraicGeometry/Pullbacks.lean b/Mathlib/AlgebraicGeometry/Pullbacks.lean index a20cda1ec5567..ca37134839d14 100644 --- a/Mathlib/AlgebraicGeometry/Pullbacks.lean +++ b/Mathlib/AlgebraicGeometry/Pullbacks.lean @@ -503,12 +503,12 @@ def openCoverOfLeftRight (𝒰X : X.OpenCover) (𝒰Y : Y.OpenCover) (f : X ⟶ def openCoverOfBase' (𝒰 : OpenCover Z) (f : X ⟶ Z) (g : Y ⟶ Z) : OpenCover (pullback f g) := by apply (openCoverOfLeft (𝒰.pullbackCover f) f g).bind intro i - have := PullbackCone.flipIsLimit <| - pasteVertIsPullback rfl (pullbackIsPullback g (𝒰.map i)) - (pullbackIsPullback (pullback.snd g (𝒰.map i)) (pullback.snd f (𝒰.map i))) + haveI := ((IsPullback.of_hasPullback (pullback.snd g (𝒰.map i)) + (pullback.snd f (𝒰.map i))).paste_horiz (IsPullback.of_hasPullback _ _)).flip refine @openCoverOfIsIso _ _ - (f := (pullbackSymmetry _ _).hom ≫ (limit.isoLimitCone ⟨_, this⟩).inv ≫ + (f := (pullbackSymmetry (pullback.snd f (𝒰.map i)) (pullback.snd g (𝒰.map i))).hom ≫ + (limit.isoLimitCone ⟨_, this.isLimit⟩).inv ≫ pullback.map _ _ _ _ (𝟙 _) (𝟙 _) (𝟙 _) ?_ ?_) inferInstance · simp [← pullback.condition] · simp only [Category.comp_id, Category.id_comp] @@ -528,12 +528,12 @@ def openCoverOfBase (𝒰 : OpenCover Z) (f : X ⟶ Z) (g : Y ⟶ Z) : OpenCover ((Equiv.prodPUnit 𝒰.J).symm.trans (Equiv.sigmaEquivProd 𝒰.J PUnit).symm) fun _ => Iso.refl _ intro i rw [Iso.refl_hom, Category.id_comp, openCoverOfBase'_map] - apply pullback.hom_ext <;> dsimp <;> - · simp only [limit.lift_π, PullbackCone.mk_pt, PullbackCone.mk_π_app, Category.assoc, - limit.lift_π_assoc, cospan_left, Category.comp_id, limit.isoLimitCone_inv_π, - limit.isoLimitCone_inv_π_assoc, PullbackCone.flip_pt, PullbackCone.π_app_left, - PullbackCone.π_app_right, PullbackCone.flip_fst, PullbackCone.flip_snd, - pullbackSymmetry_hom_comp_snd_assoc, pullbackSymmetry_hom_comp_fst_assoc] + ext : 1 <;> + · simp only [limit.lift_π, PullbackCone.mk_pt, PullbackCone.mk_π_app, Equiv.trans_apply, + Equiv.prodPUnit_symm_apply, Category.assoc, limit.lift_π_assoc, cospan_left, Category.comp_id, + limit.isoLimitCone_inv_π_assoc, PullbackCone.π_app_left, IsPullback.cone_fst, + pullbackSymmetry_hom_comp_snd_assoc, limit.isoLimitCone_inv_π, + PullbackCone.π_app_right, IsPullback.cone_snd, pullbackSymmetry_hom_comp_fst_assoc] rfl end Pullback diff --git a/Mathlib/AlgebraicGeometry/ResidueField.lean b/Mathlib/AlgebraicGeometry/ResidueField.lean index 1649e31a1b5e3..dc139a1ac5f98 100644 --- a/Mathlib/AlgebraicGeometry/ResidueField.lean +++ b/Mathlib/AlgebraicGeometry/ResidueField.lean @@ -84,19 +84,19 @@ variable {X Y : Scheme.{u}} (f : X ⟶ Y) -- We need this strange instance for `residueFieldMap`, the type of `F` must be fixed -- like this. The instance `IsLocalRingHom (f.stalkMap x)` already exists, but does not work for -- `residueFieldMap`. -instance (x): IsLocalRingHom (F := Y.presheaf.stalk (f.val.base x) →+* X.presheaf.stalk x) +instance (x): IsLocalRingHom (F := Y.presheaf.stalk (f.base x) →+* X.presheaf.stalk x) (f.stalkMap x) := - f.2 x + f.1.2 x /-- If `X ⟶ Y` is a morphism of locally ringed spaces and `x` a point of `X`, we obtain a morphism of residue fields in the other direction. -/ def Hom.residueFieldMap (f : X.Hom Y) (x : X) : - Y.residueField (f.val.base x) ⟶ X.residueField x := + Y.residueField (f.base x) ⟶ X.residueField x := LocalRing.ResidueField.map (f.stalkMap x) @[reassoc] lemma residue_residueFieldMap (x : X) : - Y.residue (f.val.base x) ≫ f.residueFieldMap x = f.stalkMap x ≫ X.residue x := by + Y.residue (f.base x) ≫ f.residueFieldMap x = f.stalkMap x ≫ X.residue x := by simp [Hom.residueFieldMap] rfl @@ -107,19 +107,19 @@ lemma residueFieldMap_id (x : X) : @[simp] lemma residueFieldMap_comp {Z : Scheme.{u}} (g : Y ⟶ Z) (x : X) : - (f ≫ g).residueFieldMap x = g.residueFieldMap (f.val.base x) ≫ f.residueFieldMap x := + (f ≫ g).residueFieldMap x = g.residueFieldMap (f.base x) ≫ f.residueFieldMap x := LocallyRingedSpace.residueFieldMap_comp _ _ _ @[reassoc] -lemma evaluation_naturality {V : Opens Y} (x : X) (hx : f.val.base x ∈ V) : - Y.evaluation V (f.val.base x) hx ≫ f.residueFieldMap x = +lemma evaluation_naturality {V : Opens Y} (x : X) (hx : f.base x ∈ V) : + Y.evaluation V (f.base x) hx ≫ f.residueFieldMap x = f.app V ≫ X.evaluation (f ⁻¹ᵁ V) x hx := - LocallyRingedSpace.evaluation_naturality f ⟨x, hx⟩ + LocallyRingedSpace.evaluation_naturality f.1 ⟨x, hx⟩ -lemma evaluation_naturality_apply {V : Opens Y} (x : X) (hx : f.val.base x ∈ V) (s) : - f.residueFieldMap x (Y.evaluation V (f.val.base x) hx s) = +lemma evaluation_naturality_apply {V : Opens Y} (x : X) (hx : f.base x ∈ V) (s) : + f.residueFieldMap x (Y.evaluation V (f.base x) hx s) = X.evaluation (f ⁻¹ᵁ V) x hx (f.app V s) := - LocallyRingedSpace.evaluation_naturality_apply f ⟨x, hx⟩ s + LocallyRingedSpace.evaluation_naturality_apply f.1 ⟨x, hx⟩ s instance [IsOpenImmersion f] (x) : IsIso (f.residueFieldMap x) := (LocalRing.ResidueField.mapEquiv diff --git a/Mathlib/AlgebraicGeometry/Restrict.lean b/Mathlib/AlgebraicGeometry/Restrict.lean index dd7a7046d4bdf..db1fe2f6dad5c 100644 --- a/Mathlib/AlgebraicGeometry/Restrict.lean +++ b/Mathlib/AlgebraicGeometry/Restrict.lean @@ -46,7 +46,7 @@ def toScheme {X : Scheme.{u}} (U : X.Opens) : Scheme.{u} := instance : CoeOut X.Opens Scheme := ⟨toScheme⟩ /-- The restriction of a scheme to an open subset. -/ -@[simps! val_base_apply] +@[simps! base_apply] def ι : ↑U ⟶ X := X.ofRestrict _ instance : IsOpenImmersion U.ι := inferInstanceAs (IsOpenImmersion (X.ofRestrict _)) @@ -81,7 +81,7 @@ lemma opensRange_ι : U.ι.opensRange = U := Opens.ext Subtype.range_val @[simp] -lemma range_ι : Set.range U.ι.val.base = U := +lemma range_ι : Set.range U.ι.base = U := Subtype.range_val lemma ι_image_top : U.ι ''ᵁ ⊤ = U := @@ -225,17 +225,17 @@ instance {U V : X.Opens} (i : U ⟶ V) : IsOpenImmersion (X.restrictFunctor.map -- Porting note: the `by ...` used to be automatically done by unification magic @[reassoc] theorem Scheme.restrictFunctor_map_ofRestrict {U V : X.Opens} (i : U ⟶ V) : - (X.restrictFunctor.map i).1 ≫ V.ι = U.ι := + (X.restrictFunctor.map i).left ≫ V.ι = U.ι := IsOpenImmersion.lift_fac _ _ (by simpa using i.le) theorem Scheme.restrictFunctor_map_base {U V : X.Opens} (i : U ⟶ V) : - (X.restrictFunctor.map i).1.val.base = (Opens.toTopCat _).map i := by + (X.restrictFunctor.map i).left.base = (Opens.toTopCat _).map i := by ext a; refine Subtype.ext ?_ -- Porting note: `ext` did not pick up `Subtype.ext` - exact (congr_arg (fun f : X.restrict U.openEmbedding ⟶ X => f.val.base a) + exact (congr_arg (fun f : X.restrict U.openEmbedding ⟶ X => f.base a) (X.restrictFunctor_map_ofRestrict i)) theorem Scheme.restrictFunctor_map_app_aux {U V : X.Opens} (i : U ⟶ V) (W : Opens V) : - U.ι ''ᵁ ((X.restrictFunctor.map i).1 ⁻¹ᵁ W) ≤ V.ι ''ᵁ W := by + U.ι ''ᵁ ((X.restrictFunctor.map i).left ⁻¹ᵁ W) ≤ V.ι ''ᵁ W := by simp only [← SetLike.coe_subset_coe, IsOpenMap.functor_obj_coe, Set.image_subset_iff, Scheme.restrictFunctor_map_base, Opens.map_coe, Opens.inclusion'_apply] rintro _ h @@ -304,7 +304,7 @@ noncomputable abbrev Scheme.restrictMapIso {X Y : Scheme.{u}} (f : X ⟶ Y) [IsI apply IsOpenImmersion.isoOfRangeEq (f := (f ⁻¹ᵁ U).ι ≫ f) U.ι _ dsimp rw [Set.range_comp, Opens.range_ι, Opens.range_ι] - refine @Set.image_preimage_eq _ _ f.val.base U.1 f.homeomorph.surjective + refine @Set.image_preimage_eq _ _ f.base U.1 f.homeomorph.surjective section MorphismRestrict @@ -374,12 +374,12 @@ instance {X Y : Scheme.{u}} (f : X ⟶ Y) [IsIso f] (U : Y.Opens) : IsIso (f ∣ delta morphismRestrict; infer_instance theorem morphismRestrict_base_coe {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) (x) : - @Coe.coe U Y (⟨fun x => x.1⟩) ((f ∣_ U).val.base x) = f.val.base x.1 := - congr_arg (fun f => PresheafedSpace.Hom.base (LocallyRingedSpace.Hom.val f) x) + @Coe.coe U Y (⟨fun x => x.1⟩) ((f ∣_ U).base x) = f.base x.1 := + congr_arg (fun f => (Scheme.Hom.toLRSHom f).base x) (morphismRestrict_ι f U) -theorem morphismRestrict_val_base {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) : - ⇑(f ∣_ U).val.base = U.1.restrictPreimage f.val.base := +theorem morphismRestrict_base {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) : + ⇑(f ∣_ U).base = U.1.restrictPreimage f.base := funext fun x => Subtype.ext (morphismRestrict_base_coe f U x) theorem image_morphismRestrict_preimage {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) (V : Opens U) : @@ -387,7 +387,7 @@ theorem image_morphismRestrict_preimage {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y. ext1 ext x constructor - · rintro ⟨⟨x, hx⟩, hx' : (f ∣_ U).val.base _ ∈ V, rfl⟩ + · rintro ⟨⟨x, hx⟩, hx' : (f ∣_ U).base _ ∈ V, rfl⟩ refine ⟨⟨_, hx⟩, ?_, rfl⟩ -- Porting note: this rewrite was not necessary rw [SetLike.mem_coe] @@ -396,7 +396,7 @@ theorem image_morphismRestrict_preimage {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y. refine Subtype.ext ?_ exact (morphismRestrict_base_coe f U ⟨x, hx⟩).symm · rintro ⟨⟨x, hx⟩, hx' : _ ∈ V.1, rfl : x = _⟩ - refine ⟨⟨_, hx⟩, (?_ : (f ∣_ U).val.base ⟨x, hx⟩ ∈ V.1), rfl⟩ + refine ⟨⟨_, hx⟩, (?_ : (f ∣_ U).base ⟨x, hx⟩ ∈ V.1), rfl⟩ convert hx' -- Porting note: `ext1` is compiling refine Subtype.ext ?_ @@ -470,7 +470,7 @@ def morphismRestrictRestrict {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) (V : (Scheme.restrictRestrict _ _ _) ?_ · ext x simp only [IsOpenMap.functor_obj_coe, Opens.coe_inclusion', - Opens.map_coe, Set.mem_image, Set.mem_preimage, SetLike.mem_coe, morphismRestrict_val_base] + Opens.map_coe, Set.mem_image, Set.mem_preimage, SetLike.mem_coe, morphismRestrict_base] constructor · rintro ⟨⟨a, h₁⟩, h₂, rfl⟩ exact ⟨_, h₂, rfl⟩ @@ -505,7 +505,7 @@ def morphismRestrictRestrictBasicOpen {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Op -/ def morphismRestrictStalkMap {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) (x) : Arrow.mk ((f ∣_ U).stalkMap x) ≅ Arrow.mk (f.stalkMap x.1) := Arrow.isoMk' _ _ - (U.stalkIso ((f ∣_ U).1.base x) ≪≫ + (U.stalkIso ((f ∣_ U).base x) ≪≫ (TopCat.Presheaf.stalkCongr _ <| Inseparable.of_eq <| morphismRestrict_base_coe f U x)) ((f ⁻¹ᵁ U).stalkIso x) <| by apply TopCat.Presheaf.stalk_hom_ext @@ -532,7 +532,7 @@ lemma resLE_eq_morphismRestrict : f.resLE U (f ⁻¹ᵁ U) le_rfl = f ∣_ U := simp [Scheme.Hom.resLE] lemma resLE_id (i : V ⟶ V') : resLE (𝟙 X) V' V i.le = (X.restrictFunctor.map i).left := by - simp only [resLE, id_val_base, morphismRestrict_id, Category.comp_id] + simp only [resLE, morphismRestrict_id] rfl @[reassoc (attr := simp)] @@ -542,7 +542,7 @@ lemma resLE_comp_ι : f.resLE U V e ≫ U.ι = V.ι ≫ f := by @[reassoc] lemma resLE_comp_resLE {Z : Scheme.{u}} (g : Y ⟶ Z) {W : Z.Opens} (e') : f.resLE U V e ≫ g.resLE W U e' = (f ≫ g).resLE W V - (e.trans ((Opens.map f.val.base).map (homOfLE e')).le) := by + (e.trans ((Opens.map f.base).map (homOfLE e')).le) := by simp [← cancel_mono W.ι] @[reassoc (attr := simp)] @@ -553,7 +553,7 @@ lemma map_resLE (i : V' ⟶ V) : @[reassoc (attr := simp)] lemma resLE_map (i : U ⟶ U') : f.resLE U V e ≫ (Y.restrictFunctor.map i).left = - f.resLE U' V (e.trans ((Opens.map f.1.base).map i).le) := by + f.resLE U' V (e.trans ((Opens.map f.base).map i).le) := by simp_rw [← resLE_id, resLE_comp_resLE, Category.comp_id] lemma resLE_congr (e₁ : U = U') (e₂ : V = V') (P : MorphismProperty Scheme.{u}) : diff --git a/Mathlib/AlgebraicGeometry/Scheme.lean b/Mathlib/AlgebraicGeometry/Scheme.lean index 7714e1f80579c..3a705c863a6a4 100644 --- a/Mathlib/AlgebraicGeometry/Scheme.lean +++ b/Mathlib/AlgebraicGeometry/Scheme.lean @@ -57,20 +57,31 @@ instance : CoeSort Scheme Type* where abbrev Opens (X : Scheme) : Type* := TopologicalSpace.Opens X /-- A morphism between schemes is a morphism between the underlying locally ringed spaces. -/ --- @[nolint has_nonempty_instance] -- Porting note(#5171): linter not ported yet -def Hom (X Y : Scheme) : Type* := - X.toLocallyRingedSpace ⟶ Y.toLocallyRingedSpace +structure Hom (X Y : Scheme) extends X.toLocallyRingedSpace.Hom Y.toLocallyRingedSpace where + +/-- Cast a morphism of schemes into morphisms of local ringed spaces. -/ +abbrev Hom.toLRSHom {X Y : Scheme.{u}} (f : X.Hom Y) : + X.toLocallyRingedSpace ⟶ Y.toLocallyRingedSpace := + f.toHom_1 + +/-- See Note [custom simps projection] -/ +def Hom.Simps.toLRSHom {X Y : Scheme.{u}} (f : X.Hom Y) : + X.toLocallyRingedSpace ⟶ Y.toLocallyRingedSpace := + f.toLRSHom + +initialize_simps_projections Hom (toHom_1 → toLRSHom) /-- Schemes are a full subcategory of locally ringed spaces. -/ -instance : Category Scheme := - { InducedCategory.category Scheme.toLocallyRingedSpace with Hom := Hom } +instance : Category Scheme where + id X := Hom.mk (𝟙 X.toLocallyRingedSpace) + comp f g := Hom.mk (f.toLRSHom ≫ g.toLRSHom) -/-- `f ⁻¹ᵁ U` is notation for `(Opens.map f.1.base).obj U`, +/-- `f ⁻¹ᵁ U` is notation for `(Opens.map f.base).obj U`, the preimage of an open set `U` under `f`. -/ scoped[AlgebraicGeometry] notation3:90 f:91 " ⁻¹ᵁ " U:90 => @Prefunctor.obj (Scheme.Opens _) _ (Scheme.Opens _) _ - (Opens.map (f : LocallyRingedSpace.Hom _ _).val.base).toPrefunctor U + (Opens.map (f : Scheme.Hom _ _).base).toPrefunctor U /-- `Γ(X, U)` is notation for `X.presheaf.obj (op U)`. -/ scoped[AlgebraicGeometry] notation3 "Γ(" X ", " U ")" => @@ -82,7 +93,7 @@ instance {X : Scheme.{u}} : Subsingleton Γ(X, ⊥) := CommRingCat.subsingleton_of_isTerminal X.sheaf.isTerminalOfEmpty @[continuity, fun_prop] -lemma Hom.continuous {X Y : Scheme} (f : X ⟶ Y) : Continuous f.1.base := f.1.base.2 +lemma Hom.continuous {X Y : Scheme} (f : X ⟶ Y) : Continuous f.base := f.base.2 /-- The structure sheaf of a scheme. -/ protected abbrev sheaf (X : Scheme) := @@ -95,12 +106,12 @@ variable {X Y : Scheme.{u}} (f : Hom X Y) {U U' : Y.Opens} {V V' : X.Opens} /-- Given a morphism of schemes `f : X ⟶ Y`, and open `U ⊆ Y`, this is the induced map `Γ(Y, U) ⟶ Γ(X, f ⁻¹ᵁ U)`. -/ abbrev app (U : Y.Opens) : Γ(Y, U) ⟶ Γ(X, f ⁻¹ᵁ U) := - f.1.c.app (op U) + f.c.app (op U) @[reassoc] lemma naturality (i : op U' ⟶ op U) : - Y.presheaf.map i ≫ f.app U = f.app U' ≫ X.presheaf.map ((Opens.map f.1.base).map i.unop).op := - f.1.c.naturality i + Y.presheaf.map i ≫ f.app U = f.app U' ≫ X.presheaf.map ((Opens.map f.base).map i.unop).op := + f.c.naturality i /-- Given a morphism of schemes `f : X ⟶ Y`, and open sets `U ⊆ Y`, `V ⊆ f ⁻¹' U`, this is the induced map `Γ(Y, U) ⟶ Γ(X, V)`. -/ @@ -121,7 +132,7 @@ lemma appLE_map' (e : V ≤ f ⁻¹ᵁ U) (i : V = V') : @[reassoc (attr := simp)] lemma map_appLE (e : V ≤ f ⁻¹ᵁ U) (i : op U' ⟶ op U) : Y.presheaf.map i ≫ f.appLE U V e = - f.appLE U' V (e.trans ((Opens.map f.1.base).map i.unop).le) := by + f.appLE U' V (e.trans ((Opens.map f.base).map i.unop).le) := by rw [Hom.appLE, f.naturality_assoc, ← Functor.map_comp] rfl @@ -143,22 +154,23 @@ lemma appLE_congr (e : V ≤ f ⁻¹ᵁ U) (e₁ : U = U') (e₂ : V = V') P (f.appLE U V e) ↔ P (f.appLE U' V' (e₁ ▸ e₂ ▸ e)) := by subst e₁; subst e₂; rfl -/-- An isomorphism of schemes induces a homeomorphism of the underlying topological spaces. -/ -noncomputable def homeomorph [IsIso f] : X ≃ₜ Y := - TopCat.homeoOfIso (asIso <| f.val.base) - -/-- A morphism of schemes `f : X ⟶ Y` induces a local ring homomorphis from `Y.presheaf.stalk (f x)` -to `X.presheaf.stalk x` for any `x : X`. -/ -def stalkMap (x : X) : Y.presheaf.stalk (f.val.base x) ⟶ X.presheaf.stalk x := - f.val.stalkMap x +/-- A morphism of schemes `f : X ⟶ Y` induces a local ring homomorphism from +`Y.presheaf.stalk (f x)` to `X.presheaf.stalk x` for any `x : X`. -/ +def stalkMap (x : X) : Y.presheaf.stalk (f.base x) ⟶ X.presheaf.stalk x := + f.toLRSHom.stalkMap x @[ext (iff := false)] -protected lemma ext {f g : X ⟶ Y} (h_base : f.1.base = g.1.base) +protected lemma ext {f g : X ⟶ Y} (h_base : f.base = g.base) (h_app : ∀ U, f.app U ≫ X.presheaf.map - (eqToHom congr((Opens.map $h_base.symm).obj U)).op = g.app U) : f = g := - LocallyRingedSpace.Hom.ext <| SheafedSpace.ext _ _ h_base + (eqToHom congr((Opens.map $h_base.symm).obj U)).op = g.app U) : f = g := by + cases f; cases g; congr 1 + exact LocallyRingedSpace.Hom.ext' <| SheafedSpace.ext _ _ h_base (TopCat.Presheaf.ext fun U ↦ by simpa using h_app U) +/-- An alternative ext lemma for scheme morphisms. -/ +protected lemma ext' {f g : X ⟶ Y} (h : f.toLRSHom = g.toLRSHom) : f = g := by + cases f; cases g; congr 1 + lemma preimage_iSup {ι} (U : ι → Opens Y) : f ⁻¹ᵁ iSup U = ⨆ i, f ⁻¹ᵁ U i := Opens.ext (by simp) @@ -177,27 +189,38 @@ lemma preimage_comp {X Y Z : Scheme.{u}} (f : X ⟶ Y) (g : Y ⟶ Z) (U) : /-- The forgetful functor from `Scheme` to `LocallyRingedSpace`. -/ @[simps!] -def forgetToLocallyRingedSpace : Scheme ⥤ LocallyRingedSpace := - inducedFunctor _ --- deriving Full, Faithful -- Porting note: no delta derive handler, see https://github.com/leanprover-community/mathlib4/issues/5020 +def forgetToLocallyRingedSpace : Scheme ⥤ LocallyRingedSpace where + obj := toLocallyRingedSpace + map := Hom.toLRSHom /-- The forget functor `Scheme ⥤ LocallyRingedSpace` is fully faithful. -/ -@[simps!] +@[simps preimage_toLRSHom] def fullyFaithfulForgetToLocallyRingedSpace : - forgetToLocallyRingedSpace.FullyFaithful := - fullyFaithfulInducedFunctor _ + forgetToLocallyRingedSpace.FullyFaithful where + preimage := Hom.mk instance : forgetToLocallyRingedSpace.Full := - InducedCategory.full _ + fullyFaithfulForgetToLocallyRingedSpace.full instance : forgetToLocallyRingedSpace.Faithful := - InducedCategory.faithful _ + fullyFaithfulForgetToLocallyRingedSpace.faithful /-- The forgetful functor from `Scheme` to `TopCat`. -/ @[simps!] def forgetToTop : Scheme ⥤ TopCat := Scheme.forgetToLocallyRingedSpace ⋙ LocallyRingedSpace.forgetToTop +/-- An isomorphism of schemes induces a homeomorphism of the underlying topological spaces. -/ +noncomputable def homeoOfIso {X Y : Scheme.{u}} (e : X ≅ Y) : X ≃ₜ Y := + TopCat.homeoOfIso (forgetToTop.mapIso e) + +alias _root_.CategoryTheory.Iso.schemeIsoToHomeo := homeoOfIso + +/-- An isomorphism of schemes induces a homeomorphism of the underlying topological spaces. -/ +noncomputable def Hom.homeomorph {X Y : Scheme.{u}} (f : X.Hom Y) [IsIso (C := Scheme) f] : + X ≃ₜ Y := + (asIso f).schemeIsoToHomeo + -- Porting note: Lean seems not able to find this coercion any more instance hasCoeToTopCat : CoeOut Scheme TopCat where coe X := X.carrier @@ -208,7 +231,7 @@ unif_hint forgetToTop_obj_eq_coe (X : Scheme) where ⊢ forgetToTop.obj X ≟ (X : TopCat) @[simp] -theorem id_val_base (X : Scheme) : (𝟙 X : _).1.base = 𝟙 _ := +theorem id.base (X : Scheme) : (𝟙 X : _).base = 𝟙 _ := rfl @[simp] @@ -216,22 +239,23 @@ theorem id_app {X : Scheme} (U : X.Opens) : (𝟙 X : _).app U = 𝟙 _ := rfl @[reassoc] -theorem comp_val {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) : (f ≫ g).val = f.val ≫ g.val := +theorem comp_toLRSHom {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) : + (f ≫ g).toLRSHom = f.toLRSHom ≫ g.toLRSHom := rfl @[simp, reassoc] -- reassoc lemma does not need `simp` theorem comp_coeBase {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) : - (f ≫ g).val.base = f.val.base ≫ g.val.base := + (f ≫ g).base = f.base ≫ g.base := rfl -- Porting note: removed elementwise attribute, as generated lemmas were trivial. @[reassoc] -theorem comp_val_base {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) : - (f ≫ g).val.base = f.val.base ≫ g.val.base := +theorem comp_base {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) : + (f ≫ g).base = f.base ≫ g.base := rfl -theorem comp_val_base_apply {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) (x : X) : - (f ≫ g).val.base x = g.val.base (f.val.base x) := by +theorem comp_base_apply {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) (x : X) : + (f ≫ g).base x = g.base (f.base x) := by simp @[simp, reassoc] -- reassoc lemma does not need `simp` @@ -244,7 +268,7 @@ theorem comp_app {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) (U) : theorem appLE_comp_appLE {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) (U V W e₁ e₂) : g.appLE U V e₁ ≫ f.appLE V W e₂ = - (f ≫ g).appLE U W (e₂.trans ((Opens.map f.1.base).map (homOfLE e₁)).le) := by + (f ≫ g).appLE U W (e₂.trans ((Opens.map f.base).map (homOfLE e₁)).le) := by dsimp [Hom.appLE] rw [Category.assoc, f.naturality_assoc, ← Functor.map_comp] rfl @@ -262,8 +286,8 @@ theorem app_eq {X Y : Scheme} (f : X ⟶ Y) {U V : Y.Opens} (e : U = V) : f.app U = Y.presheaf.map (eqToHom e.symm).op ≫ f.app V ≫ - X.presheaf.map (eqToHom (congr_arg (Opens.map f.val.base).obj e)).op := by - rw [← IsIso.inv_comp_eq, ← Functor.map_inv, f.val.c.naturality] + X.presheaf.map (eqToHom (congr_arg (Opens.map f.base).obj e)).op := by + rw [← IsIso.inv_comp_eq, ← Functor.map_inv, f.naturality] cases e rfl @@ -276,21 +300,20 @@ lemma presheaf_map_eqToHom_op (X : Scheme) (U V : X.Opens) (i : U = V) : X.presheaf.map (eqToHom i).op = eqToHom (i ▸ rfl) := by rw [eqToHom_op, eqToHom_map] -instance is_locallyRingedSpace_iso {X Y : Scheme} (f : X ⟶ Y) [IsIso f] : - @IsIso LocallyRingedSpace _ _ _ f := +instance is_locallyRingedSpace_iso {X Y : Scheme} (f : X ⟶ Y) [IsIso f] : IsIso f.toLRSHom := forgetToLocallyRingedSpace.map_isIso f -instance val_base_isIso {X Y : Scheme.{u}} (f : X ⟶ Y) [IsIso f] : IsIso f.1.base := +instance base_isIso {X Y : Scheme.{u}} (f : X ⟶ Y) [IsIso f] : IsIso f.base := Scheme.forgetToTop.map_isIso f -- Porting note: need an extra instance here. -instance {X Y : Scheme} (f : X ⟶ Y) [IsIso f] (U) : IsIso (f.val.c.app U) := - haveI := PresheafedSpace.c_isIso_of_iso f.val - NatIso.isIso_app_of_isIso _ _ +instance {X Y : Scheme} (f : X ⟶ Y) [IsIso f] (U) : IsIso (f.c.app U) := + haveI := PresheafedSpace.c_isIso_of_iso f.toPshHom + NatIso.isIso_app_of_isIso f.c _ instance {X Y : Scheme} (f : X ⟶ Y) [IsIso f] (U) : IsIso (f.app U) := - haveI := PresheafedSpace.c_isIso_of_iso f.val - NatIso.isIso_app_of_isIso _ _ + haveI := PresheafedSpace.c_isIso_of_iso f.toPshHom + NatIso.isIso_app_of_isIso f.c _ @[simp] theorem inv_app {X Y : Scheme} (f : X ⟶ Y) [IsIso f] (U : X.Opens) : @@ -318,16 +341,16 @@ theorem Spec_toLocallyRingedSpace (R : CommRingCat) : /-- The induced map of a ring homomorphism on the ring spectra, as a morphism of schemes. -/ def Spec.map {R S : CommRingCat} (f : R ⟶ S) : Spec S ⟶ Spec R := - (Spec.locallyRingedSpaceMap f : Spec.locallyRingedSpaceObj S ⟶ Spec.locallyRingedSpaceObj R) + ⟨Spec.locallyRingedSpaceMap f⟩ @[simp] theorem Spec.map_id (R : CommRingCat) : Spec.map (𝟙 R) = 𝟙 (Spec R) := - Spec.locallyRingedSpaceMap_id R + Scheme.Hom.ext' <| Spec.locallyRingedSpaceMap_id R @[reassoc, simp] theorem Spec.map_comp {R S T : CommRingCat} (f : R ⟶ S) (g : S ⟶ T) : Spec.map (f ≫ g) = Spec.map g ≫ Spec.map f := - Spec.locallyRingedSpaceMap_comp f g + Scheme.Hom.ext' <| Spec.locallyRingedSpaceMap_comp f g /-- The spectrum, as a contravariant functor from commutative rings to schemes. -/ @[simps] @@ -358,7 +381,7 @@ variable {R S : CommRingCat.{u}} (f : R ⟶ S) lemma Spec_carrier (R : CommRingCat.{u}) : (Spec R).carrier = PrimeSpectrum R := rfl lemma Spec_sheaf (R : CommRingCat.{u}) : (Spec R).sheaf = Spec.structureSheaf R := rfl lemma Spec_presheaf (R : CommRingCat.{u}) : (Spec R).presheaf = (Spec.structureSheaf R).1 := rfl -lemma Spec.map_base : (Spec.map f).1.base = PrimeSpectrum.comap f := rfl +lemma Spec.map_base : (Spec.map f).base = PrimeSpectrum.comap f := rfl lemma Spec.map_app (U) : (Spec.map f).app U = StructureSheaf.comap f U (Spec.map f ⁻¹ᵁ U) le_rfl := rfl @@ -388,9 +411,9 @@ instance : Inhabited Scheme := /-- The global sections, notated Gamma. -/ def Γ : Schemeᵒᵖ ⥤ CommRingCat := - (inducedFunctor Scheme.toLocallyRingedSpace).op ⋙ LocallyRingedSpace.Γ + Scheme.forgetToLocallyRingedSpace.op ⋙ LocallyRingedSpace.Γ -theorem Γ_def : Γ = (inducedFunctor Scheme.toLocallyRingedSpace).op ⋙ LocallyRingedSpace.Γ := +theorem Γ_def : Γ = Scheme.forgetToLocallyRingedSpace.op ⋙ LocallyRingedSpace.Γ := rfl @[simp] @@ -491,7 +514,7 @@ lemma basicOpen_restrict (i : V ⟶ U) (f : Γ(X, U)) : @[simp] theorem preimage_basicOpen {X Y : Scheme.{u}} (f : X ⟶ Y) {U : Y.Opens} (r : Γ(Y, U)) : f ⁻¹ᵁ (Y.basicOpen r) = X.basicOpen (f.app U r) := - LocallyRingedSpace.preimage_basicOpen f r + LocallyRingedSpace.preimage_basicOpen f.toLRSHom r lemma basicOpen_appLE {X Y : Scheme.{u}} (f : X ⟶ Y) (U : X.Opens) (V : Y.Opens) (e : U ≤ f ⁻¹ᵁ V) (s : Γ(Y, V)) : X.basicOpen (f.appLE V U e s) = U ⊓ f ⁻¹ᵁ (Y.basicOpen s) := by @@ -582,25 +605,25 @@ theorem Scheme.Spec_map_presheaf_map_eqToHom {X : Scheme} {U V : X.Opens} (h : U simp [eqToHom_map] @[reassoc (attr := simp)] -lemma Scheme.iso_hom_val_base_inv_val_base {X Y : Scheme.{u}} (e : X ≅ Y) : - e.hom.val.base ≫ e.inv.val.base = 𝟙 _ := - LocallyRingedSpace.iso_hom_val_base_inv_val_base (Scheme.forgetToLocallyRingedSpace.mapIso e) +lemma Scheme.iso_hom_base_inv_base {X Y : Scheme.{u}} (e : X ≅ Y) : + e.hom.base ≫ e.inv.base = 𝟙 _ := + LocallyRingedSpace.iso_hom_base_inv_base (Scheme.forgetToLocallyRingedSpace.mapIso e) @[simp] -lemma Scheme.iso_hom_val_base_inv_val_base_apply {X Y : Scheme.{u}} (e : X ≅ Y) (x : X) : - (e.inv.val.base (e.hom.val.base x)) = x := by - show (e.hom.val.base ≫ e.inv.val.base) x = 𝟙 X.toPresheafedSpace x +lemma Scheme.iso_hom_base_inv_base_apply {X Y : Scheme.{u}} (e : X ≅ Y) (x : X) : + (e.inv.base (e.hom.base x)) = x := by + show (e.hom.base ≫ e.inv.base) x = 𝟙 X.toPresheafedSpace x simp @[reassoc (attr := simp)] -lemma Scheme.iso_inv_val_base_hom_val_base {X Y : Scheme.{u}} (e : X ≅ Y) : - e.inv.val.base ≫ e.hom.val.base = 𝟙 _ := - LocallyRingedSpace.iso_inv_val_base_hom_val_base (Scheme.forgetToLocallyRingedSpace.mapIso e) +lemma Scheme.iso_inv_base_hom_base {X Y : Scheme.{u}} (e : X ≅ Y) : + e.inv.base ≫ e.hom.base = 𝟙 _ := + LocallyRingedSpace.iso_inv_base_hom_base (Scheme.forgetToLocallyRingedSpace.mapIso e) @[simp] -lemma Scheme.iso_inv_val_base_hom_val_base_apply {X Y : Scheme.{u}} (e : X ≅ Y) (y : Y) : - (e.hom.val.base (e.inv.val.base y)) = y := by - show (e.inv.val.base ≫ e.hom.val.base) y = 𝟙 Y.toPresheafedSpace y +lemma Scheme.iso_inv_base_hom_base_apply {X Y : Scheme.{u}} (e : X ≅ Y) (y : Y) : + (e.hom.base (e.inv.base y)) = y := by + show (e.inv.base ≫ e.hom.base) y = 𝟙 Y.toPresheafedSpace y simp section Stalks @@ -615,17 +638,17 @@ lemma stalkMap_id (X : Scheme.{u}) (x : X) : PresheafedSpace.stalkMap.id _ x lemma stalkMap_comp {X Y Z : Scheme.{u}} (f : X ⟶ Y) (g : Y ⟶ Z) (x : X) : - (f ≫ g : X ⟶ Z).stalkMap x = g.stalkMap (f.val.base x) ≫ f.stalkMap x := - PresheafedSpace.stalkMap.comp f.val g.val x + (f ≫ g : X ⟶ Z).stalkMap x = g.stalkMap (f.base x) ≫ f.stalkMap x := + PresheafedSpace.stalkMap.comp f.toPshHom g.toPshHom x @[reassoc] lemma stalkSpecializes_stalkMap (x x' : X) - (h : x ⤳ x') : Y.presheaf.stalkSpecializes (f.val.base.map_specializes h) ≫ f.stalkMap x = + (h : x ⤳ x') : Y.presheaf.stalkSpecializes (f.base.map_specializes h) ≫ f.stalkMap x = f.stalkMap x' ≫ X.presheaf.stalkSpecializes h := - PresheafedSpace.stalkMap.stalkSpecializes_stalkMap f.val h + PresheafedSpace.stalkMap.stalkSpecializes_stalkMap f.toPshHom h lemma stalkSpecializes_stalkMap_apply (x x' : X) (h : x ⤳ x') (y) : - f.stalkMap x (Y.presheaf.stalkSpecializes (f.val.base.map_specializes h) y) = + f.stalkMap x (Y.presheaf.stalkSpecializes (f.base.map_specializes h) y) = (X.presheaf.stalkSpecializes h (f.stalkMap x' y)) := DFunLike.congr_fun (stalkSpecializes_stalkMap f x x' h) y @@ -633,54 +656,54 @@ lemma stalkSpecializes_stalkMap_apply (x x' : X) (h : x ⤳ x') (y) : lemma stalkMap_congr (f g : X ⟶ Y) (hfg : f = g) (x x' : X) (hxx' : x = x') : f.stalkMap x ≫ (X.presheaf.stalkCongr (.of_eq hxx')).hom = (Y.presheaf.stalkCongr (.of_eq <| hfg ▸ hxx' ▸ rfl)).hom ≫ g.stalkMap x' := - LocallyRingedSpace.stalkMap_congr f g hfg x x' hxx' + LocallyRingedSpace.stalkMap_congr f.toLRSHom g.toLRSHom congr(($hfg).toLRSHom) x x' hxx' @[reassoc] lemma stalkMap_congr_hom (f g : X ⟶ Y) (hfg : f = g) (x : X) : f.stalkMap x = (Y.presheaf.stalkCongr (.of_eq <| hfg ▸ rfl)).hom ≫ g.stalkMap x := - LocallyRingedSpace.stalkMap_congr_hom f g hfg x + LocallyRingedSpace.stalkMap_congr_hom f.toLRSHom g.toLRSHom congr(($hfg).toLRSHom) x @[reassoc] lemma stalkMap_congr_point (x x' : X) (hxx' : x = x') : f.stalkMap x ≫ (X.presheaf.stalkCongr (.of_eq hxx')).hom = (Y.presheaf.stalkCongr (.of_eq <| hxx' ▸ rfl)).hom ≫ f.stalkMap x' := - LocallyRingedSpace.stalkMap_congr_point f x x' hxx' + LocallyRingedSpace.stalkMap_congr_point f.toLRSHom x x' hxx' @[reassoc (attr := simp)] lemma stalkMap_hom_inv (e : X ≅ Y) (y : Y) : - e.hom.stalkMap (e.inv.val.base y) ≫ e.inv.stalkMap y = + e.hom.stalkMap (e.inv.base y) ≫ e.inv.stalkMap y = (Y.presheaf.stalkCongr (.of_eq (by simp))).hom := LocallyRingedSpace.stalkMap_hom_inv (forgetToLocallyRingedSpace.mapIso e) y @[simp] lemma stalkMap_hom_inv_apply (e : X ≅ Y) (y : Y) (z) : - e.inv.stalkMap y (e.hom.stalkMap (e.inv.val.base y) z) = + e.inv.stalkMap y (e.hom.stalkMap (e.inv.base y) z) = (Y.presheaf.stalkCongr (.of_eq (by simp))).hom z := DFunLike.congr_fun (stalkMap_hom_inv e y) z @[reassoc (attr := simp)] lemma stalkMap_inv_hom (e : X ≅ Y) (x : X) : - e.inv.stalkMap (e.hom.val.base x) ≫ e.hom.stalkMap x = + e.inv.stalkMap (e.hom.base x) ≫ e.hom.stalkMap x = (X.presheaf.stalkCongr (.of_eq (by simp))).hom := LocallyRingedSpace.stalkMap_inv_hom (forgetToLocallyRingedSpace.mapIso e) x @[simp] lemma stalkMap_inv_hom_apply (e : X ≅ Y) (x : X) (y) : - e.hom.stalkMap x (e.inv.stalkMap (e.hom.val.base x) y) = + e.hom.stalkMap x (e.inv.stalkMap (e.hom.base x) y) = (X.presheaf.stalkCongr (.of_eq (by simp))).hom y := DFunLike.congr_fun (stalkMap_inv_hom e x) y @[reassoc (attr := simp)] -lemma stalkMap_germ (U : Y.Opens) (x : X) (hx : f.val.base x ∈ U) : - Y.presheaf.germ U (f.val.base x) hx ≫ f.stalkMap x = +lemma stalkMap_germ (U : Y.Opens) (x : X) (hx : f.base x ∈ U) : + Y.presheaf.germ U (f.base x) hx ≫ f.stalkMap x = f.app U ≫ X.presheaf.germ (f ⁻¹ᵁ U) x hx := - PresheafedSpace.stalkMap_germ f.val U x hx + PresheafedSpace.stalkMap_germ f.toPshHom U x hx @[simp] -lemma stalkMap_germ_apply (U : Y.Opens) (x : X) (hx : f.val.base x ∈ U) (y) : - f.stalkMap x (Y.presheaf.germ _ (f.val.base x) hx y) = +lemma stalkMap_germ_apply (U : Y.Opens) (x : X) (hx : f.base x ∈ U) (y) : + f.stalkMap x (Y.presheaf.germ _ (f.base x) hx y) = X.presheaf.germ (f ⁻¹ᵁ U) x hx (f.app U y) := - PresheafedSpace.stalkMap_germ_apply f.val U x hx y + PresheafedSpace.stalkMap_germ_apply f.toPshHom U x hx y end Scheme diff --git a/Mathlib/AlgebraicGeometry/Spec.lean b/Mathlib/AlgebraicGeometry/Spec.lean index 80690a03c2022..349fe5db9ccec 100644 --- a/Mathlib/AlgebraicGeometry/Spec.lean +++ b/Mathlib/AlgebraicGeometry/Spec.lean @@ -226,7 +226,7 @@ theorem localRingHom_comp_stalkIso {R S : CommRingCat.{u}} (f : R ⟶ S) (p : Pr /-- The induced map of a ring homomorphism on the prime spectra, as a morphism of locally ringed spaces. -/ -@[simps] +@[simps toShHom] def Spec.locallyRingedSpaceMap {R S : CommRingCat.{u}} (f : R ⟶ S) : Spec.locallyRingedSpaceObj S ⟶ Spec.locallyRingedSpaceObj R := LocallyRingedSpace.Hom.mk (Spec.sheafedSpaceMap f) fun p => @@ -251,14 +251,14 @@ def Spec.locallyRingedSpaceMap {R S : CommRingCat.{u}} (f : R ⟶ S) : @[simp] theorem Spec.locallyRingedSpaceMap_id (R : CommRingCat.{u}) : Spec.locallyRingedSpaceMap (𝟙 R) = 𝟙 (Spec.locallyRingedSpaceObj R) := - LocallyRingedSpace.Hom.ext <| by - rw [Spec.locallyRingedSpaceMap_val, Spec.sheafedSpaceMap_id]; rfl + LocallyRingedSpace.Hom.ext' <| by + rw [Spec.locallyRingedSpaceMap_toShHom, Spec.sheafedSpaceMap_id]; rfl theorem Spec.locallyRingedSpaceMap_comp {R S T : CommRingCat.{u}} (f : R ⟶ S) (g : S ⟶ T) : Spec.locallyRingedSpaceMap (f ≫ g) = Spec.locallyRingedSpaceMap g ≫ Spec.locallyRingedSpaceMap f := - LocallyRingedSpace.Hom.ext <| by - rw [Spec.locallyRingedSpaceMap_val, Spec.sheafedSpaceMap_comp]; rfl + LocallyRingedSpace.Hom.ext' <| by + rw [Spec.locallyRingedSpaceMap_toShHom, Spec.sheafedSpaceMap_comp]; rfl /-- Spec, as a contravariant functor from commutative rings to locally ringed spaces. -/ diff --git a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace.lean b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace.lean index 0f98d08e80581..72d5eb095aa96 100644 --- a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace.lean +++ b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace.lean @@ -70,44 +70,59 @@ def 𝒪 : Sheaf CommRingCat X.toTopCat := /-- A morphism of locally ringed spaces is a morphism of ringed spaces such that the morphisms induced on stalks are local ring homomorphisms. -/ @[ext] -structure Hom (X Y : LocallyRingedSpace.{u}) : Type _ where - /-- the underlying morphism between ringed spaces -/ - val : X.toSheafedSpace ⟶ Y.toSheafedSpace +structure Hom (X Y : LocallyRingedSpace.{u}) + extends X.toPresheafedSpace.Hom Y.toPresheafedSpace : Type _ where /-- the underlying morphism induces a local ring homomorphism on stalks -/ - prop : ∀ x, IsLocalRingHom (val.stalkMap x) + prop : ∀ x, IsLocalRingHom (toHom.stalkMap x) + +/-- A morphism of locally ringed spaces as a morphism of sheafed spaces. -/ +abbrev Hom.toShHom {X Y : LocallyRingedSpace.{u}} (f : X.Hom Y) : + X.toSheafedSpace ⟶ Y.toSheafedSpace := f.1 + +@[simp, nolint simpVarHead] +lemma Hom.toShHom_mk {X Y : LocallyRingedSpace.{u}} + (f : X.toPresheafedSpace.Hom Y.toPresheafedSpace) (hf) : + Hom.toShHom ⟨f, hf⟩ = f := rfl instance : Quiver LocallyRingedSpace := ⟨Hom⟩ -@[ext] lemma Hom.ext' (X Y : LocallyRingedSpace.{u}) {f g : X ⟶ Y} (h : f.val = g.val) : f = g := - Hom.ext h +@[ext] lemma Hom.ext' {X Y : LocallyRingedSpace.{u}} {f g : X ⟶ Y} (h : f.toShHom = g.toShHom) : + f = g := by cases f; cases g; congr + +/-- See Note [custom simps projection] -/ +def Hom.Simps.toShHom {X Y : LocallyRingedSpace.{u}} (f : X.Hom Y) : + X.toSheafedSpace ⟶ Y.toSheafedSpace := + f.toShHom + +initialize_simps_projections Hom (toHom → toShHom) /-- A morphism of locally ringed spaces `f : X ⟶ Y` induces a local ring homomorphism from `Y.stalk (f x)` to `X.stalk x` for any `x : X`. -/ noncomputable def Hom.stalkMap {X Y : LocallyRingedSpace.{u}} (f : Hom X Y) (x : X) : Y.presheaf.stalk (f.1.1 x) ⟶ X.presheaf.stalk x := - f.val.stalkMap x + f.toShHom.stalkMap x instance isLocalRingHomStalkMap {X Y : LocallyRingedSpace.{u}} (f : X ⟶ Y) (x : X) : IsLocalRingHom (f.stalkMap x) := f.2 x instance isLocalRingHomValStalkMap {X Y : LocallyRingedSpace.{u}} (f : X ⟶ Y) (x : X) : - IsLocalRingHom (f.val.stalkMap x) := + IsLocalRingHom (f.toShHom.stalkMap x) := f.2 x /-- The identity morphism on a locally ringed space. -/ -@[simps] +@[simps! toShHom] def id (X : LocallyRingedSpace.{u}) : Hom X X := - ⟨𝟙 _, fun x => by erw [PresheafedSpace.stalkMap.id]; apply isLocalRingHom_id⟩ + ⟨𝟙 X.toSheafedSpace, fun x => by erw [PresheafedSpace.stalkMap.id]; apply isLocalRingHom_id⟩ instance (X : LocallyRingedSpace.{u}) : Inhabited (Hom X X) := ⟨id X⟩ /-- Composition of morphisms of locally ringed spaces. -/ def comp {X Y Z : LocallyRingedSpace.{u}} (f : Hom X Y) (g : Hom Y Z) : Hom X Z := - ⟨f.val ≫ g.val, fun x => by + ⟨f.toShHom ≫ g.toShHom, fun x => by erw [PresheafedSpace.stalkMap.comp] exact @RingHom.isLocalRingHom_comp _ _ _ _ _ _ _ _ (f.2 _) (g.2 _)⟩ @@ -116,9 +131,9 @@ instance : Category LocallyRingedSpace.{u} where Hom := Hom id := id comp {_ _ _} f g := comp f g - comp_id {X Y} f := Hom.ext <| by simp [comp] - id_comp {X Y} f := Hom.ext <| by simp [comp] - assoc {_ _ _ _} f g h := Hom.ext <| by simp [comp] + comp_id {X Y} f := Hom.ext' <| by simp [comp] + id_comp {X Y} f := Hom.ext' <| by simp [comp] + assoc {_ _ _ _} f g h := Hom.ext' <| by simp [comp] /-- The forgetful functor from `LocallyRingedSpace` to `SheafedSpace CommRing`. -/ @[simps] @@ -126,8 +141,9 @@ def forgetToSheafedSpace : LocallyRingedSpace.{u} ⥤ SheafedSpace CommRingCat.{ obj X := X.toSheafedSpace map {_ _} f := f.1 +/-- The canonical map `X ⟶ Spec Γ(X, ⊤)`. This is the unit of the `Γ-Spec` adjunction. -/ instance : forgetToSheafedSpace.Faithful where - map_injective {_ _} _ _ h := Hom.ext h + map_injective {_ _} _ _ h := Hom.ext' h /-- The forgetful functor from `LocallyRingedSpace` to `Top`. -/ @[simps!] @@ -135,22 +151,25 @@ def forgetToTop : LocallyRingedSpace.{u} ⥤ TopCat.{u} := forgetToSheafedSpace ⋙ SheafedSpace.forget _ @[simp] -theorem comp_val {X Y Z : LocallyRingedSpace.{u}} (f : X ⟶ Y) (g : Y ⟶ Z) : - (f ≫ g).val = f.val ≫ g.val := +theorem comp_toShHom {X Y Z : LocallyRingedSpace.{u}} (f : X ⟶ Y) (g : Y ⟶ Z) : + (f ≫ g).toShHom = f.toShHom ≫ g.toShHom := rfl -@[simp] theorem id_val' (X : LocallyRingedSpace.{u}) : Hom.val (𝟙 X) = 𝟙 X.toSheafedSpace := +/-- A variant of `id_toShHom'` that works with `𝟙 X` instead of `id X`. -/ +@[simp] theorem id_toShHom' (X : LocallyRingedSpace.{u}) : + Hom.toShHom (𝟙 X) = 𝟙 X.toSheafedSpace := rfl --- Porting note: complains that `(f ≫ g).val.c` can be further simplified --- so changed to its simp normal form `(f.val ≫ g.val).c` -@[simp] -theorem comp_val_c {X Y Z : LocallyRingedSpace.{u}} (f : X ⟶ Y) (g : Y ⟶ Z) : - (f.1 ≫ g.1).c = g.val.c ≫ (Presheaf.pushforward _ g.val.base).map f.val.c := +theorem comp_base {X Y Z : LocallyRingedSpace.{u}} (f : X ⟶ Y) (g : Y ⟶ Z) : + (f ≫ g).base = f.base ≫ g.base := + rfl + +theorem comp_c {X Y Z : LocallyRingedSpace.{u}} (f : X ⟶ Y) (g : Y ⟶ Z) : + (f ≫ g).c = g.c ≫ (Presheaf.pushforward _ g.base).map f.c := rfl -theorem comp_val_c_app {X Y Z : LocallyRingedSpace.{u}} (f : X ⟶ Y) (g : Y ⟶ Z) (U : (Opens Z)ᵒᵖ) : - (f ≫ g).val.c.app U = g.val.c.app U ≫ f.val.c.app (op <| (Opens.map g.val.base).obj U.unop) := +theorem comp_c_app {X Y Z : LocallyRingedSpace.{u}} (f : X ⟶ Y) (g : Y ⟶ Z) (U : (Opens Z)ᵒᵖ) : + (f ≫ g).c.app U = g.c.app U ≫ f.c.app (op <| (Opens.map g.base).obj U.unop) := rfl /-- Given two locally ringed spaces `X` and `Y`, an isomorphism between `X` and `Y` as _sheafed_ @@ -158,7 +177,7 @@ spaces can be lifted to a morphism `X ⟶ Y` as locally ringed spaces. See also `isoOfSheafedSpaceIso`. -/ -@[simps] +@[simps! toShHom] def homOfSheafedSpaceHomOfIsIso {X Y : LocallyRingedSpace.{u}} (f : X.toSheafedSpace ⟶ Y.toSheafedSpace) [IsIso f] : X ⟶ Y := Hom.mk f fun x => @@ -179,15 +198,16 @@ def isoOfSheafedSpaceIso {X Y : LocallyRingedSpace.{u}} (f : X.toSheafedSpace X ≅ Y where hom := homOfSheafedSpaceHomOfIsIso f.hom inv := homOfSheafedSpaceHomOfIsIso f.inv - hom_inv_id := Hom.ext f.hom_inv_id - inv_hom_id := Hom.ext f.inv_hom_id + hom_inv_id := Hom.ext' f.hom_inv_id + inv_hom_id := Hom.ext' f.inv_hom_id instance : forgetToSheafedSpace.ReflectsIsomorphisms where reflects {_ _} f i := { out := ⟨homOfSheafedSpaceHomOfIsIso (CategoryTheory.inv (forgetToSheafedSpace.map f)), - Hom.ext (IsIso.hom_inv_id (I := i)), Hom.ext (IsIso.inv_hom_id (I := i))⟩ } + Hom.ext' (IsIso.hom_inv_id (I := i)), Hom.ext' (IsIso.inv_hom_id (I := i))⟩ } -instance is_sheafedSpace_iso {X Y : LocallyRingedSpace.{u}} (f : X ⟶ Y) [IsIso f] : IsIso f.1 := +instance is_sheafedSpace_iso {X Y : LocallyRingedSpace.{u}} (f : X ⟶ Y) [IsIso f] : + IsIso f.toShHom := LocallyRingedSpace.forgetToSheafedSpace.map_isIso f /-- The restriction of a locally ringed space along an open embedding. @@ -229,10 +249,10 @@ theorem Γ_obj_op (X : LocallyRingedSpace.{u}) : Γ.obj (op X) = X.presheaf.obj rfl @[simp] -theorem Γ_map {X Y : LocallyRingedSpace.{u}ᵒᵖ} (f : X ⟶ Y) : Γ.map f = f.unop.1.c.app (op ⊤) := +theorem Γ_map {X Y : LocallyRingedSpace.{u}ᵒᵖ} (f : X ⟶ Y) : Γ.map f = f.unop.c.app (op ⊤) := rfl -theorem Γ_map_op {X Y : LocallyRingedSpace.{u}} (f : X ⟶ Y) : Γ.map f.op = f.1.c.app (op ⊤) := +theorem Γ_map_op {X Y : LocallyRingedSpace.{u}} (f : X ⟶ Y) : Γ.map f.op = f.c.app (op ⊤) := rfl /-- The empty locally ringed space. -/ @@ -289,27 +309,27 @@ instance component_nontrivial (X : LocallyRingedSpace.{u}) (U : Opens X.carrier) (X.presheaf.germ _ _ hU.some.2).domain_nontrivial @[simp] -lemma iso_hom_val_base_inv_val_base {X Y : LocallyRingedSpace.{u}} (e : X ≅ Y) : - e.hom.val.base ≫ e.inv.val.base = 𝟙 _ := by - rw [← SheafedSpace.comp_base, ← LocallyRingedSpace.comp_val] +lemma iso_hom_base_inv_base {X Y : LocallyRingedSpace.{u}} (e : X ≅ Y) : + e.hom.base ≫ e.inv.base = 𝟙 _ := by + rw [← SheafedSpace.comp_base, ← LocallyRingedSpace.comp_toShHom] simp @[simp] -lemma iso_hom_val_base_inv_val_base_apply {X Y : LocallyRingedSpace.{u}} (e : X ≅ Y) (x : X) : - (e.inv.val.base (e.hom.val.base x)) = x := by - show (e.hom.val.base ≫ e.inv.val.base) x = 𝟙 X.toPresheafedSpace x +lemma iso_hom_base_inv_base_apply {X Y : LocallyRingedSpace.{u}} (e : X ≅ Y) (x : X) : + (e.inv.base (e.hom.base x)) = x := by + show (e.hom.base ≫ e.inv.base) x = 𝟙 X.toPresheafedSpace x simp @[simp] -lemma iso_inv_val_base_hom_val_base {X Y : LocallyRingedSpace.{u}} (e : X ≅ Y) : - e.inv.val.base ≫ e.hom.val.base = 𝟙 _ := by - rw [← SheafedSpace.comp_base, ← LocallyRingedSpace.comp_val] +lemma iso_inv_base_hom_base {X Y : LocallyRingedSpace.{u}} (e : X ≅ Y) : + e.inv.base ≫ e.hom.base = 𝟙 _ := by + rw [← SheafedSpace.comp_base, ← LocallyRingedSpace.comp_toShHom] simp @[simp] -lemma iso_inv_val_base_hom_val_base_apply {X Y : LocallyRingedSpace.{u}} (e : X ≅ Y) (y : Y) : - (e.hom.val.base (e.inv.val.base y)) = y := by - show (e.inv.val.base ≫ e.hom.val.base) y = 𝟙 Y.toPresheafedSpace y +lemma iso_inv_base_hom_base_apply {X Y : LocallyRingedSpace.{u}} (e : X ≅ Y) (y : Y) : + (e.hom.base (e.inv.base y)) = y := by + show (e.inv.base ≫ e.hom.base) y = 𝟙 Y.toPresheafedSpace y simp section Stalks @@ -322,17 +342,17 @@ lemma stalkMap_id (X : LocallyRingedSpace.{u}) (x : X) : PresheafedSpace.stalkMap.id _ x lemma stalkMap_comp (x : X) : - (f ≫ g : X ⟶ Z).stalkMap x = g.stalkMap (f.val.base x) ≫ f.stalkMap x := - PresheafedSpace.stalkMap.comp f.val g.val x + (f ≫ g : X ⟶ Z).stalkMap x = g.stalkMap (f.base x) ≫ f.stalkMap x := + PresheafedSpace.stalkMap.comp f.toShHom g.toShHom x @[reassoc] lemma stalkSpecializes_stalkMap (x x' : X) (h : x ⤳ x') : - Y.presheaf.stalkSpecializes (f.val.base.map_specializes h) ≫ f.stalkMap x = + Y.presheaf.stalkSpecializes (f.base.map_specializes h) ≫ f.stalkMap x = f.stalkMap x' ≫ X.presheaf.stalkSpecializes h := - PresheafedSpace.stalkMap.stalkSpecializes_stalkMap f.val h + PresheafedSpace.stalkMap.stalkSpecializes_stalkMap f.toShHom h lemma stalkSpecializes_stalkMap_apply (x x' : X) (h : x ⤳ x') (y) : - f.stalkMap x (Y.presheaf.stalkSpecializes (f.val.base.map_specializes h) y) = + f.stalkMap x (Y.presheaf.stalkSpecializes (f.base.map_specializes h) y) = (X.presheaf.stalkSpecializes h (f.stalkMap x' y)) := DFunLike.congr_fun (stalkSpecializes_stalkMap f x x' h) y @@ -360,57 +380,57 @@ lemma stalkMap_congr_point {X Y : LocallyRingedSpace.{u}} (f : X ⟶ Y) (x x' : @[reassoc (attr := simp)] lemma stalkMap_hom_inv (e : X ≅ Y) (y : Y) : - e.hom.stalkMap (e.inv.val.base y) ≫ e.inv.stalkMap y = + e.hom.stalkMap (e.inv.base y) ≫ e.inv.stalkMap y = Y.presheaf.stalkSpecializes (specializes_of_eq <| by simp) := by rw [← stalkMap_comp, LocallyRingedSpace.stalkMap_congr_hom (e.inv ≫ e.hom) (𝟙 _) (by simp)] simp @[simp] lemma stalkMap_hom_inv_apply (e : X ≅ Y) (y : Y) (z) : - e.inv.stalkMap y (e.hom.stalkMap (e.inv.val.base y) z) = + e.inv.stalkMap y (e.hom.stalkMap (e.inv.base y) z) = Y.presheaf.stalkSpecializes (specializes_of_eq <| by simp) z := DFunLike.congr_fun (stalkMap_hom_inv e y) z @[reassoc (attr := simp)] lemma stalkMap_inv_hom (e : X ≅ Y) (x : X) : - e.inv.stalkMap (e.hom.val.base x) ≫ e.hom.stalkMap x = + e.inv.stalkMap (e.hom.base x) ≫ e.hom.stalkMap x = X.presheaf.stalkSpecializes (specializes_of_eq <| by simp) := by rw [← stalkMap_comp, LocallyRingedSpace.stalkMap_congr_hom (e.hom ≫ e.inv) (𝟙 _) (by simp)] simp @[simp] lemma stalkMap_inv_hom_apply (e : X ≅ Y) (x : X) (y) : - e.hom.stalkMap x (e.inv.stalkMap (e.hom.val.base x) y) = + e.hom.stalkMap x (e.inv.stalkMap (e.hom.base x) y) = X.presheaf.stalkSpecializes (specializes_of_eq <| by simp) y := DFunLike.congr_fun (stalkMap_inv_hom e x) y @[reassoc (attr := simp)] -lemma stalkMap_germ (U : Opens Y) (x : X) (hx : f.val.base x ∈ U) : - Y.presheaf.germ U (f.val.base x) hx ≫ f.stalkMap x = - f.val.c.app (op U) ≫ X.presheaf.germ ((Opens.map f.1.base).obj U) x hx := - PresheafedSpace.stalkMap_germ f.val U x hx +lemma stalkMap_germ (U : Opens Y) (x : X) (hx : f.base x ∈ U) : + Y.presheaf.germ U (f.base x) hx ≫ f.stalkMap x = + f.c.app (op U) ≫ X.presheaf.germ ((Opens.map f.base).obj U) x hx := + PresheafedSpace.stalkMap_germ f.toShHom U x hx -lemma stalkMap_germ_apply (U : Opens Y) (x : X) (hx : f.val.base x ∈ U) (y) : - f.stalkMap x (Y.presheaf.germ U (f.val.base x) hx y) = - X.presheaf.germ ((Opens.map f.1.base).obj U) x hx (f.val.c.app (op U) y) := - PresheafedSpace.stalkMap_germ_apply f.val U x hx y +lemma stalkMap_germ_apply (U : Opens Y) (x : X) (hx : f.base x ∈ U) (y) : + f.stalkMap x (Y.presheaf.germ U (f.base x) hx y) = + X.presheaf.germ ((Opens.map f.base).obj U) x hx (f.c.app (op U) y) := + PresheafedSpace.stalkMap_germ_apply f.toShHom U x hx y theorem preimage_basicOpen {X Y : LocallyRingedSpace.{u}} (f : X ⟶ Y) {U : Opens Y} (s : Y.presheaf.obj (op U)) : - (Opens.map f.1.base).obj (Y.toRingedSpace.basicOpen s) = - @RingedSpace.basicOpen X.toRingedSpace ((Opens.map f.1.base).obj U) (f.1.c.app _ s) := by + (Opens.map f.base).obj (Y.toRingedSpace.basicOpen s) = + @RingedSpace.basicOpen X.toRingedSpace ((Opens.map f.base).obj U) (f.c.app _ s) := by ext x constructor · rintro ⟨hxU, hx⟩ rw [SetLike.mem_coe, X.toRingedSpace.mem_basicOpen _ _ hxU] delta toRingedSpace rw [← stalkMap_germ_apply] - exact (f.val.stalkMap _).isUnit_map hx + exact (f.stalkMap _).isUnit_map hx · rintro ⟨hxU, hx⟩ simp only [Opens.map_coe, Set.mem_preimage, SetLike.mem_coe, toRingedSpace] at hx ⊢ - rw [RingedSpace.mem_basicOpen _ s (f.1.base x) hxU] + rw [RingedSpace.mem_basicOpen _ s (f.base x) hxU] rw [← stalkMap_germ_apply] at hx - exact (isUnit_map_iff (f.val.stalkMap _) _).mp hx + exact (isUnit_map_iff (f.toShHom.stalkMap _) _).mp hx variable {U : TopCat} (X : LocallyRingedSpace.{u}) {f : U ⟶ X.toTopCat} (h : OpenEmbedding f) (V : Opens U) (x : U) (hx : x ∈ V) diff --git a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean index fac6ae0294b93..a45f070a1bf4f 100644 --- a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean +++ b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean @@ -106,12 +106,12 @@ noncomputable def coproductCofanIsColimit : IsColimit (coproductCofan F) where (((forgetToSheafedSpace.mapCocone s).ι.app i).stalkMap y) := (s.ι.app i).2 y infer_instance⟩ - fac _ _ := LocallyRingedSpace.Hom.ext + fac _ _ := LocallyRingedSpace.Hom.ext' (colimit.ι_desc (C := SheafedSpace.{u+1, u, u} CommRingCatMax.{u, u}) _ _) uniq s f h := - LocallyRingedSpace.Hom.ext - (IsColimit.uniq _ (forgetToSheafedSpace.mapCocone s) f.1 fun j => - congr_arg LocallyRingedSpace.Hom.val (h j)) + LocallyRingedSpace.Hom.ext' + (IsColimit.uniq _ (forgetToSheafedSpace.mapCocone s) f.toShHom fun j => + congr_arg LocallyRingedSpace.Hom.toShHom (h j)) instance : HasCoproducts.{u} LocallyRingedSpace.{u} := fun _ => ⟨fun F => ⟨⟨⟨_, coproductCofanIsColimit F⟩⟩⟩⟩ @@ -132,14 +132,15 @@ variable {X Y : LocallyRingedSpace.{v}} (f g : X ⟶ Y) namespace HasCoequalizer instance coequalizer_π_app_isLocalRingHom - (U : TopologicalSpace.Opens (coequalizer f.val g.val).carrier) : - IsLocalRingHom ((coequalizer.π f.val g.val : _).c.app (op U)) := by - have := ι_comp_coequalizerComparison f.1 g.1 SheafedSpace.forgetToPresheafedSpace + (U : TopologicalSpace.Opens (coequalizer f.toShHom g.toShHom).carrier) : + IsLocalRingHom ((coequalizer.π f.toShHom g.toShHom : _).c.app (op U)) := by + have := ι_comp_coequalizerComparison f.toShHom g.toShHom SheafedSpace.forgetToPresheafedSpace rw [← PreservesCoequalizer.iso_hom] at this erw [SheafedSpace.congr_app this.symm (op U)] rw [PresheafedSpace.comp_c_app, ← PresheafedSpace.colimitPresheafObjIsoComponentwiseLimit_hom_π] -- Porting note (#10754): this instance has to be manually added - haveI : IsIso (PreservesCoequalizer.iso SheafedSpace.forgetToPresheafedSpace f.val g.val).hom.c := + haveI : IsIso (PreservesCoequalizer.iso + SheafedSpace.forgetToPresheafedSpace f.toShHom g.toShHom).hom.c := PresheafedSpace.c_isIso_of_iso _ infer_instance @@ -162,46 +163,46 @@ are local ring homs. -/ -variable (U : Opens (coequalizer f.1 g.1).carrier) -variable (s : (coequalizer f.1 g.1).presheaf.obj (op U)) +variable (U : Opens (coequalizer f.toShHom g.toShHom).carrier) +variable (s : (coequalizer f.toShHom g.toShHom).presheaf.obj (op U)) /-- (Implementation). The basic open set of the section `π꙳ s`. -/ noncomputable def imageBasicOpen : Opens Y := Y.toRingedSpace.basicOpen - (show Y.presheaf.obj (op (unop _)) from ((coequalizer.π f.1 g.1).c.app (op U)) s) + (show Y.presheaf.obj (op (unop _)) from ((coequalizer.π f.toShHom g.toShHom).c.app (op U)) s) theorem imageBasicOpen_image_preimage : - (coequalizer.π f.1 g.1).base ⁻¹' ((coequalizer.π f.1 g.1).base '' (imageBasicOpen f g U s).1) = - (imageBasicOpen f g U s).1 := by - fapply Types.coequalizer_preimage_image_eq_of_preimage_eq - -- Porting note: Type of `f.1.base` and `g.1.base` needs to be explicit - (f.1.base : X.carrier.1 ⟶ Y.carrier.1) (g.1.base : X.carrier.1 ⟶ Y.carrier.1) + (coequalizer.π f.toShHom g.toShHom).base ⁻¹' ((coequalizer.π f.toShHom g.toShHom).base '' + (imageBasicOpen f g U s).1) = (imageBasicOpen f g U s).1 := by + fapply Types.coequalizer_preimage_image_eq_of_preimage_eq f.base + -- Porting note: Type of `f.base` and `g.base` needs to be explicit + (g.base : X.carrier.1 ⟶ Y.carrier.1) · ext simp_rw [types_comp_apply, ← TopCat.comp_app, ← PresheafedSpace.comp_base] congr 2 - exact coequalizer.condition f.1 g.1 + exact coequalizer.condition f.toShHom g.toShHom · apply isColimitCoforkMapOfIsColimit (forget TopCat) apply isColimitCoforkMapOfIsColimit (SheafedSpace.forget _) - exact coequalizerIsCoequalizer f.1 g.1 + exact coequalizerIsCoequalizer f.toShHom g.toShHom · suffices - (TopologicalSpace.Opens.map f.1.base).obj (imageBasicOpen f g U s) = - (TopologicalSpace.Opens.map g.1.base).obj (imageBasicOpen f g U s) + (TopologicalSpace.Opens.map f.base).obj (imageBasicOpen f g U s) = + (TopologicalSpace.Opens.map g.base).obj (imageBasicOpen f g U s) by injection this delta imageBasicOpen rw [preimage_basicOpen f, preimage_basicOpen g] dsimp only [Functor.op, unop_op] -- Porting note (#11224): change `rw` to `erw` erw [← comp_apply, ← SheafedSpace.comp_c_app', ← comp_apply, ← SheafedSpace.comp_c_app', - SheafedSpace.congr_app (coequalizer.condition f.1 g.1), comp_apply, + SheafedSpace.congr_app (coequalizer.condition f.toShHom g.toShHom), comp_apply, X.toRingedSpace.basicOpen_res] apply inf_eq_right.mpr refine (RingedSpace.basicOpen_le _ _).trans ?_ - rw [coequalizer.condition f.1 g.1] + rw [coequalizer.condition f.toShHom g.toShHom] theorem imageBasicOpen_image_open : - IsOpen ((coequalizer.π f.1 g.1).base '' (imageBasicOpen f g U s).1) := by - rw [← (TopCat.homeoOfIso (PreservesCoequalizer.iso (SheafedSpace.forget _) f.1 - g.1)).isOpen_preimage, TopCat.coequalizer_isOpen_iff, ← Set.preimage_comp] + IsOpen ((coequalizer.π f.toShHom g.toShHom).base '' (imageBasicOpen f g U s).1) := by + rw [← (TopCat.homeoOfIso (PreservesCoequalizer.iso (SheafedSpace.forget _) f.toShHom + g.toShHom)).isOpen_preimage, TopCat.coequalizer_isOpen_iff, ← Set.preimage_comp] erw [← TopCat.coe_comp] rw [PreservesCoequalizer.iso_hom, ι_comp_coequalizerComparison] dsimp only [SheafedSpace.forget] @@ -210,52 +211,54 @@ theorem imageBasicOpen_image_open : exact (imageBasicOpen f g U s).2 instance coequalizer_π_stalk_isLocalRingHom (x : Y) : - IsLocalRingHom ((coequalizer.π f.val g.val : _).stalkMap x) := by + IsLocalRingHom ((coequalizer.π f.toShHom g.toShHom : _).stalkMap x) := by constructor rintro a ha rcases TopCat.Presheaf.germ_exist _ _ a with ⟨U, hU, s, rfl⟩ - rw [PresheafedSpace.stalkMap_germ_apply (coequalizer.π f.1 g.1 : _) U _ hU] at ha + rw [PresheafedSpace.stalkMap_germ_apply + (coequalizer.π (C := SheafedSpace _) f.toShHom g.toShHom) U _ hU] at ha let V := imageBasicOpen f g U s - have hV : (coequalizer.π f.1 g.1).base ⁻¹' ((coequalizer.π f.1 g.1).base '' V.1) = V.1 := + have hV : (coequalizer.π f.toShHom g.toShHom).base ⁻¹' + ((coequalizer.π f.toShHom g.toShHom).base '' V.1) = V.1 := imageBasicOpen_image_preimage f g U s - have hV' : - V = ⟨(coequalizer.π f.1 g.1).base ⁻¹' ((coequalizer.π f.1 g.1).base '' V.1), hV.symm ▸ V.2⟩ := + have hV' : V = ⟨(coequalizer.π f.toShHom g.toShHom).base ⁻¹' + ((coequalizer.π f.toShHom g.toShHom).base '' V.1), hV.symm ▸ V.2⟩ := SetLike.ext' hV.symm - have V_open : IsOpen ((coequalizer.π f.val g.val).base '' V.1) := + have V_open : IsOpen ((coequalizer.π f.toShHom g.toShHom).base '' V.1) := imageBasicOpen_image_open f g U s - have VleU : (⟨(coequalizer.π f.val g.val).base '' V.1, V_open⟩ : TopologicalSpace.Opens _) ≤ U := + have VleU : (⟨(coequalizer.π f.toShHom g.toShHom).base '' V.1, V_open⟩ : _) ≤ U := Set.image_subset_iff.mpr (Y.toRingedSpace.basicOpen_le _) have hxV : x ∈ V := ⟨hU, ha⟩ rw [← - (coequalizer f.val g.val).presheaf.germ_res_apply (homOfLE VleU) _ - (@Set.mem_image_of_mem _ _ (coequalizer.π f.val g.val).base x V.1 hxV) s] + (coequalizer f.toShHom g.toShHom).presheaf.germ_res_apply (homOfLE VleU) _ + (@Set.mem_image_of_mem _ _ (coequalizer.π f.toShHom g.toShHom).base x V.1 hxV) s] apply RingHom.isUnit_map - rw [← isUnit_map_iff ((coequalizer.π f.val g.val : _).c.app _), ← comp_apply, + rw [← isUnit_map_iff ((coequalizer.π f.toShHom g.toShHom : _).c.app _), ← comp_apply, NatTrans.naturality, comp_apply, ← isUnit_map_iff (Y.presheaf.map (eqToHom hV').op)] -- Porting note (#11224): change `rw` to `erw` erw [← comp_apply, ← comp_apply, ← Y.presheaf.map_comp] convert @RingedSpace.isUnit_res_basicOpen Y.toRingedSpace (unop _) - (((coequalizer.π f.val g.val).c.app (op U)) s) + (((coequalizer.π f.toShHom g.toShHom).c.app (op U)) s) end HasCoequalizer /-- The coequalizer of two locally ringed space in the category of sheafed spaces is a locally ringed space. -/ noncomputable def coequalizer : LocallyRingedSpace where - toSheafedSpace := Limits.coequalizer f.1 g.1 + toSheafedSpace := Limits.coequalizer f.toShHom g.toShHom localRing x := by obtain ⟨y, rfl⟩ := - (TopCat.epi_iff_surjective (coequalizer.π f.val g.val).base).mp inferInstance x + (TopCat.epi_iff_surjective (coequalizer.π f.toShHom g.toShHom).base).mp inferInstance x -- TODO: this instance was found automatically before #6045 - have _ : IsLocalRingHom ((coequalizer.π f.val g.val).stalkMap y) := inferInstance - exact ((coequalizer.π f.val g.val : _).stalkMap y).domain_localRing + have _ : IsLocalRingHom ((coequalizer.π f.toShHom g.toShHom).stalkMap y) := inferInstance + exact ((coequalizer.π f.toShHom g.toShHom : _).stalkMap y).domain_localRing /-- The explicit coequalizer cofork of locally ringed spaces. -/ noncomputable def coequalizerCofork : Cofork f g := - @Cofork.ofπ _ _ _ _ f g (coequalizer f g) ⟨coequalizer.π f.1 g.1, + @Cofork.ofπ _ _ _ _ f g (coequalizer f g) ⟨coequalizer.π f.toShHom g.toShHom, -- Porting note: this used to be automatic HasCoequalizer.coequalizer_π_stalk_isLocalRingHom _ _⟩ - (LocallyRingedSpace.Hom.ext (coequalizer.condition f.1 g.1)) + (LocallyRingedSpace.Hom.ext' (coequalizer.condition f.toShHom g.toShHom)) theorem isLocalRingHom_stalkMap_congr {X Y : RingedSpace} (f g : X ⟶ Y) (H : f = g) (x) (h : IsLocalRingHom (f.stalkMap x)) : @@ -266,11 +269,11 @@ theorem isLocalRingHom_stalkMap_congr {X Y : RingedSpace} (f g : X ⟶ Y) (H : f noncomputable def coequalizerCoforkIsColimit : IsColimit (coequalizerCofork f g) := by apply Cofork.IsColimit.mk' intro s - have e : f.val ≫ s.π.val = g.val ≫ s.π.val := by injection s.condition - refine ⟨⟨coequalizer.desc s.π.1 e, ?_⟩, ?_⟩ + have e : f.toShHom ≫ s.π.toShHom = g.toShHom ≫ s.π.toShHom := by injection s.condition + refine ⟨⟨coequalizer.desc s.π.toShHom e, ?_⟩, ?_⟩ · intro x - rcases (TopCat.epi_iff_surjective (coequalizer.π f.val g.val).base).mp inferInstance x with - ⟨y, rfl⟩ + rcases (TopCat.epi_iff_surjective + (coequalizer.π f.toShHom g.toShHom).base).mp inferInstance x with ⟨y, rfl⟩ -- Porting note: was `apply isLocalRingHom_of_comp _ (PresheafedSpace.stalkMap ...)`, this -- used to allow you to provide the proof that `... ≫ ...` is a local ring homomorphism later, -- but this is no longer possible @@ -281,18 +284,18 @@ noncomputable def coequalizerCoforkIsColimit : IsColimit (coequalizerCofork f g) -- note to reviewers: this `change` is now more brittle because it now has to fully resolve -- the type to be able to search for `MonoidHomClass`, even though of course all homs in -- `CommRingCat` are clearly such - change IsLocalRingHom (h ≫ (coequalizerCofork f g).π.val.stalkMap y) + change IsLocalRingHom (h ≫ (coequalizerCofork f g).π.toShHom.stalkMap y) erw [← PresheafedSpace.stalkMap.comp] - apply isLocalRingHom_stalkMap_congr _ _ (coequalizer.π_desc s.π.1 e).symm y + apply isLocalRingHom_stalkMap_congr _ _ (coequalizer.π_desc s.π.toShHom e).symm y infer_instance constructor - · exact LocallyRingedSpace.Hom.ext (coequalizer.π_desc _ _) + · exact LocallyRingedSpace.Hom.ext' (coequalizer.π_desc _ _) intro m h - replace h : (coequalizerCofork f g).π.1 ≫ m.1 = s.π.1 := by rw [← h]; rfl - apply LocallyRingedSpace.Hom.ext - apply (colimit.isColimit (parallelPair f.1 g.1)).uniq (Cofork.ofπ s.π.1 e) m.1 + replace h : (coequalizerCofork f g).π.toShHom ≫ m.1 = s.π.toShHom := by rw [← h]; rfl + apply LocallyRingedSpace.Hom.ext' + apply (colimit.isColimit (parallelPair f.toShHom g.toShHom)).uniq (Cofork.ofπ s.π.toShHom e) m.1 rintro ⟨⟩ - · rw [← (colimit.cocone (parallelPair f.val g.val)).w WalkingParallelPairHom.left, + · rw [← (colimit.cocone (parallelPair f.toShHom g.toShHom)).w WalkingParallelPairHom.left, Category.assoc] change _ ≫ _ ≫ _ = _ ≫ _ congr diff --git a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/ResidueField.lean b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/ResidueField.lean index 44d7e9d141d92..b09ad64961538 100644 --- a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/ResidueField.lean +++ b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/ResidueField.lean @@ -83,13 +83,13 @@ variable {X Y : LocallyRingedSpace.{u}} (f : X ⟶ Y) (x : X) -- We need this strange instance for `residueFieldMap`, the type of `F` must be fixed -- like this. The instance `IsLocalRingHom (f.stalkMap x)` already exists, but does not work for -- `residueFieldMap`. -instance : IsLocalRingHom (F := Y.presheaf.stalk (f.val.base x) →+* X.presheaf.stalk x) +instance : IsLocalRingHom (F := Y.presheaf.stalk (f.base x) →+* X.presheaf.stalk x) (f.stalkMap x) := f.2 x /-- If `X ⟶ Y` is a morphism of locally ringed spaces and `x` a point of `X`, we obtain a morphism of residue fields in the other direction. -/ -def residueFieldMap (x : X) : Y.residueField (f.val.base x) ⟶ X.residueField x := +def residueFieldMap (x : X) : Y.residueField (f.base x) ⟶ X.residueField x := LocalRing.ResidueField.map (f.stalkMap x) lemma residue_comp_residueFieldMap_eq_stalkMap_comp_residue (x : X) : @@ -100,23 +100,23 @@ lemma residue_comp_residueFieldMap_eq_stalkMap_comp_residue (x : X) : @[simp] lemma residueFieldMap_id (x : X) : residueFieldMap (𝟙 X) x = 𝟙 (X.residueField x) := by - simp only [id_val', SheafedSpace.id_base, TopCat.coe_id, id_eq, residueFieldMap, stalkMap_id] + simp only [id_toShHom', SheafedSpace.id_base, TopCat.coe_id, id_eq, residueFieldMap, stalkMap_id] apply LocalRing.ResidueField.map_id @[simp] lemma residueFieldMap_comp {Z : LocallyRingedSpace.{u}} (g : Y ⟶ Z) (x : X) : - residueFieldMap (f ≫ g) x = residueFieldMap g (f.val.base x) ≫ residueFieldMap f x := by - simp only [comp_val, SheafedSpace.comp_base, Function.comp_apply, residueFieldMap] + residueFieldMap (f ≫ g) x = residueFieldMap g (f.base x) ≫ residueFieldMap f x := by + simp only [comp_toShHom, SheafedSpace.comp_base, Function.comp_apply, residueFieldMap] simp_rw [stalkMap_comp] - haveI : IsLocalRingHom (g.stalkMap (f.val.base x)) := inferInstance + haveI : IsLocalRingHom (g.stalkMap (f.base x)) := inferInstance -- TODO: This instance is found before #6045. haveI : IsLocalRingHom (f.stalkMap x) := inferInstance apply LocalRing.ResidueField.map_comp @[reassoc] -lemma evaluation_naturality {V : Opens Y} (x : (Opens.map f.1.base).obj V) : - Y.evaluation ⟨f.val.base x, x.property⟩ ≫ residueFieldMap f x.val = - f.val.c.app (op V) ≫ X.evaluation x := by +lemma evaluation_naturality {V : Opens Y} (x : (Opens.map f.base).obj V) : + Y.evaluation ⟨f.base x, x.property⟩ ≫ residueFieldMap f x.val = + f.c.app (op V) ≫ X.evaluation x := by dsimp only [LocallyRingedSpace.evaluation, LocallyRingedSpace.residueFieldMap] rw [Category.assoc] @@ -125,21 +125,21 @@ lemma evaluation_naturality {V : Opens Y} (x : (Opens.map f.1.base).obj V) : erw [LocalRing.ResidueField.map_residue, PresheafedSpace.stalkMap_germ_apply] rfl -lemma evaluation_naturality_apply {V : Opens Y} (x : (Opens.map f.1.base).obj V) +lemma evaluation_naturality_apply {V : Opens Y} (x : (Opens.map f.base).obj V) (a : Y.presheaf.obj (op V)) : - residueFieldMap f x.val (Y.evaluation ⟨f.val.base x, x.property⟩ a) = - X.evaluation x (f.val.c.app (op V) a) := by + residueFieldMap f x.val (Y.evaluation ⟨f.base x, x.property⟩ a) = + X.evaluation x (f.c.app (op V) a) := by simpa using congrFun (congrArg DFunLike.coe <| evaluation_naturality f x) a @[reassoc] lemma Γevaluation_naturality (x : X) : - Y.Γevaluation (f.val.base x) ≫ residueFieldMap f x = - f.val.c.app (op ⊤) ≫ X.Γevaluation x := + Y.Γevaluation (f.base x) ≫ residueFieldMap f x = + f.c.app (op ⊤) ≫ X.Γevaluation x := evaluation_naturality f ⟨x, by simp only [Opens.map_top]; trivial⟩ lemma Γevaluation_naturality_apply (x : X) (a : Y.presheaf.obj (op ⊤)) : - residueFieldMap f x (Y.Γevaluation (f.val.base x) a) = - X.Γevaluation x (f.val.c.app (op ⊤) a) := + residueFieldMap f x (Y.Γevaluation (f.base x) a) = + X.Γevaluation x (f.c.app (op ⊤) a) := evaluation_naturality_apply f ⟨x, by simp only [Opens.map_top]; trivial⟩ a end LocallyRingedSpace diff --git a/Mathlib/Geometry/RingedSpace/OpenImmersion.lean b/Mathlib/Geometry/RingedSpace/OpenImmersion.lean index a5571115fc47d..0f4f4ce207163 100644 --- a/Mathlib/Geometry/RingedSpace/OpenImmersion.lean +++ b/Mathlib/Geometry/RingedSpace/OpenImmersion.lean @@ -577,7 +577,7 @@ def toLocallyRingedSpaceHom : toLocallyRingedSpace Y f ⟶ Y := ⟨f, fun _ => inferInstance⟩ @[simp] -theorem toLocallyRingedSpaceHom_val : (toLocallyRingedSpaceHom Y f).val = f := +theorem toLocallyRingedSpaceHom_val : (toLocallyRingedSpaceHom Y f).toShHom = f := rfl instance toLocallyRingedSpace_isOpenImmersion : @@ -933,15 +933,15 @@ variable [H : LocallyRingedSpace.IsOpenImmersion f] instance (priority := 100) of_isIso [IsIso g] : LocallyRingedSpace.IsOpenImmersion g := @PresheafedSpace.IsOpenImmersion.ofIsIso _ _ _ _ g.1 ⟨⟨(inv g).1, by - erw [← LocallyRingedSpace.comp_val]; rw [IsIso.hom_inv_id] - erw [← LocallyRingedSpace.comp_val]; rw [IsIso.inv_hom_id]; constructor <;> rfl⟩⟩ + erw [← LocallyRingedSpace.comp_toShHom]; rw [IsIso.hom_inv_id] + erw [← LocallyRingedSpace.comp_toShHom]; rw [IsIso.inv_hom_id]; constructor <;> rfl⟩⟩ instance comp (g : Z ⟶ Y) [LocallyRingedSpace.IsOpenImmersion g] : LocallyRingedSpace.IsOpenImmersion (f ≫ g) := PresheafedSpace.IsOpenImmersion.comp f.1 g.1 instance mono : Mono f := - LocallyRingedSpace.forgetToSheafedSpace.mono_of_mono_map (show Mono f.1 by infer_instance) + LocallyRingedSpace.forgetToSheafedSpace.mono_of_mono_map (show Mono f.toShHom by infer_instance) instance : SheafedSpace.IsOpenImmersion (LocallyRingedSpace.forgetToSheafedSpace.map f) := H @@ -951,7 +951,7 @@ set_option synthInstance.maxHeartbeats 30000 in /-- An explicit pullback cone over `cospan f g` if `f` is an open immersion. -/ def pullbackConeOfLeft : PullbackCone f g := by refine PullbackCone.mk ?_ - (Y.ofRestrict (TopCat.snd_openEmbedding_of_left_openEmbedding H.base_open g.1.base)) ?_ + (Y.ofRestrict (TopCat.snd_openEmbedding_of_left_openEmbedding H.base_open g.base)) ?_ · use PresheafedSpace.IsOpenImmersion.pullbackConeOfLeftFst f.1 g.1 intro x have := PresheafedSpace.stalkMap.congr_hom _ _ @@ -960,7 +960,7 @@ def pullbackConeOfLeft : PullbackCone f g := by rw [← IsIso.eq_inv_comp] at this rw [this] infer_instance - · exact LocallyRingedSpace.Hom.ext + · exact LocallyRingedSpace.Hom.ext' (PresheafedSpace.IsOpenImmersion.pullback_cone_of_left_condition _ _) instance : LocallyRingedSpace.IsOpenImmersion (pullbackConeOfLeft f g).snd := @@ -971,16 +971,17 @@ set_option synthInstance.maxHeartbeats 80000 in def pullbackConeOfLeftIsLimit : IsLimit (pullbackConeOfLeft f g) := PullbackCone.isLimitAux' _ fun s => by refine ⟨LocallyRingedSpace.Hom.mk (PresheafedSpace.IsOpenImmersion.pullbackConeOfLeftLift - f.1 g.1 (PullbackCone.mk _ _ (congr_arg LocallyRingedSpace.Hom.val s.condition))) ?_, - LocallyRingedSpace.Hom.ext + f.1 g.1 (PullbackCone.mk _ _ (congr_arg LocallyRingedSpace.Hom.toShHom s.condition))) ?_, + LocallyRingedSpace.Hom.ext' (PresheafedSpace.IsOpenImmersion.pullbackConeOfLeftLift_fst f.1 g.1 _), - LocallyRingedSpace.Hom.ext + LocallyRingedSpace.Hom.ext' (PresheafedSpace.IsOpenImmersion.pullbackConeOfLeftLift_snd f.1 g.1 _), ?_⟩ · intro x have := PresheafedSpace.stalkMap.congr_hom _ _ (PresheafedSpace.IsOpenImmersion.pullbackConeOfLeftLift_snd f.1 g.1 - (PullbackCone.mk s.fst.1 s.snd.1 (congr_arg LocallyRingedSpace.Hom.val s.condition))) + (PullbackCone.mk s.fst.1 s.snd.1 + (congr_arg LocallyRingedSpace.Hom.toShHom s.condition))) x change _ = _ ≫ s.snd.1.stalkMap x at this rw [PresheafedSpace.stalkMap.comp, ← IsIso.eq_inv_comp] at this @@ -989,9 +990,10 @@ def pullbackConeOfLeftIsLimit : IsLimit (pullbackConeOfLeft f g) := apply CommRingCat.isLocalRingHom_comp · intro m _ h₂ rw [← cancel_mono (pullbackConeOfLeft f g).snd] - exact h₂.trans <| LocallyRingedSpace.Hom.ext + exact h₂.trans <| LocallyRingedSpace.Hom.ext' (PresheafedSpace.IsOpenImmersion.pullbackConeOfLeftLift_snd f.1 g.1 <| - PullbackCone.mk s.fst.1 s.snd.1 <| congr_arg LocallyRingedSpace.Hom.val s.condition).symm + PullbackCone.mk s.fst.1 s.snd.1 <| congr_arg + LocallyRingedSpace.Hom.toShHom s.condition).symm instance hasPullback_of_left : HasPullback f g := ⟨⟨⟨_, pullbackConeOfLeftIsLimit f g⟩⟩⟩ @@ -1082,7 +1084,7 @@ instance forgetToPresheafedSpaceReflectsPullbackOfRight : (LocallyRingedSpace.forgetToSheafedSpace ⋙ SheafedSpace.forgetToPresheafedSpace) := reflectsLimitOfReflectsIsomorphisms _ _ -theorem pullback_snd_isIso_of_range_subset (H' : Set.range g.1.base ⊆ Set.range f.1.base) : +theorem pullback_snd_isIso_of_range_subset (H' : Set.range g.base ⊆ Set.range f.base) : IsIso (pullback.snd f g) := by apply (config := {allowSynthFailures := true}) Functor.ReflectsIsomorphisms.reflects (F := LocallyRingedSpace.forgetToSheafedSpace) @@ -1099,35 +1101,35 @@ For an open immersion `f : X ⟶ Z`, given any morphism of schemes `g : Y ⟶ Z` image is contained in the image of `f`, we can lift this morphism to a unique `Y ⟶ X` that commutes with these maps. -/ -def lift (H' : Set.range g.1.base ⊆ Set.range f.1.base) : Y ⟶ X := +def lift (H' : Set.range g.base ⊆ Set.range f.base) : Y ⟶ X := -- Porting note (#10754): added instance manually have := pullback_snd_isIso_of_range_subset f g H' inv (pullback.snd f g) ≫ pullback.fst _ _ @[simp, reassoc] -theorem lift_fac (H' : Set.range g.1.base ⊆ Set.range f.1.base) : lift f g H' ≫ f = g := by +theorem lift_fac (H' : Set.range g.base ⊆ Set.range f.base) : lift f g H' ≫ f = g := by erw [Category.assoc]; rw [IsIso.inv_comp_eq]; exact pullback.condition -theorem lift_uniq (H' : Set.range g.1.base ⊆ Set.range f.1.base) (l : Y ⟶ X) (hl : l ≫ f = g) : +theorem lift_uniq (H' : Set.range g.base ⊆ Set.range f.base) (l : Y ⟶ X) (hl : l ≫ f = g) : l = lift f g H' := by rw [← cancel_mono f, hl, lift_fac] -theorem lift_range (H' : Set.range g.1.base ⊆ Set.range f.1.base) : - Set.range (lift f g H').1.base = f.1.base ⁻¹' Set.range g.1.base := by +theorem lift_range (H' : Set.range g.base ⊆ Set.range f.base) : + Set.range (lift f g H').base = f.base ⁻¹' Set.range g.base := by -- Porting note (#10754): added instance manually have := pullback_snd_isIso_of_range_subset f g H' dsimp only [lift] - have : _ = (pullback.fst f g).val.base := + have : _ = (pullback.fst f g).base := PreservesPullback.iso_hom_fst (LocallyRingedSpace.forgetToSheafedSpace ⋙ SheafedSpace.forget _) f g - rw [LocallyRingedSpace.comp_val, SheafedSpace.comp_base, ← this, ← Category.assoc, coe_comp, - Set.range_comp, Set.range_iff_surjective.mpr, Set.image_univ] + rw [LocallyRingedSpace.comp_base, ← this, ← Category.assoc, coe_comp] + rw [Set.range_comp, Set.range_iff_surjective.mpr, Set.image_univ] -- Porting note (#11224): change `rw` to `erw` on this lemma · erw [TopCat.pullback_fst_range] ext constructor · rintro ⟨y, eq⟩; exact ⟨y, eq.symm⟩ · rintro ⟨y, eq⟩; exact ⟨y, eq.symm⟩ - · rw [← TopCat.epi_iff_surjective, show (inv (pullback.snd f g)).val.base = _ from + · rw [← TopCat.epi_iff_surjective, show (inv (pullback.snd f g)).base = _ from (LocallyRingedSpace.forgetToSheafedSpace ⋙ SheafedSpace.forget _).map_inv _] infer_instance @@ -1153,7 +1155,7 @@ whose forgetful functor reflects isomorphisms, preserves limits and filtered col Then a morphism `X ⟶ Y` that is a topological open embedding is an open immersion iff every stalk map is an iso. -/ -theorem of_stalk_iso {X Y : LocallyRingedSpace} (f : X ⟶ Y) (hf : OpenEmbedding f.1.base) +theorem of_stalk_iso {X Y : LocallyRingedSpace} (f : X ⟶ Y) (hf : OpenEmbedding f.base) [stalk_iso : ∀ x : X.1, IsIso (f.stalkMap x)] : LocallyRingedSpace.IsOpenImmersion f := SheafedSpace.IsOpenImmersion.of_stalk_iso _ hf (H := stalk_iso) @@ -1190,13 +1192,13 @@ instance (U : Opens X) : IsIso (H.invApp _ U) := by delta invApp; infer_instance theorem inv_invApp (U : Opens X) : inv (H.invApp _ U) = - f.1.c.app (op (opensFunctor f |>.obj U)) ≫ + f.c.app (op (opensFunctor f |>.obj U)) ≫ X.presheaf.map (eqToHom (by simp [Opens.map, Set.preimage_image_eq _ H.base_open.inj])) := PresheafedSpace.IsOpenImmersion.inv_invApp f.1 U @[reassoc (attr := simp)] theorem invApp_app (U : Opens X) : - H.invApp _ U ≫ f.1.c.app (op (opensFunctor f |>.obj U)) = + H.invApp _ U ≫ f.c.app (op (opensFunctor f |>.obj U)) = X.presheaf.map (eqToHom (by simp [Opens.map, Set.preimage_image_eq _ H.base_open.inj])) := PresheafedSpace.IsOpenImmersion.invApp_app f.1 U @@ -1204,20 +1206,20 @@ attribute [elementwise] invApp_app @[reassoc (attr := simp)] theorem app_invApp (U : Opens Y) : - f.1.c.app (op U) ≫ H.invApp _ ((Opens.map f.1.base).obj U) = + f.c.app (op U) ≫ H.invApp _ ((Opens.map f.base).obj U) = Y.presheaf.map - ((homOfLE (Set.image_preimage_subset f.1.base U.1)).op : - op U ⟶ op (opensFunctor f |>.obj ((Opens.map f.1.base).obj U))) := + ((homOfLE (Set.image_preimage_subset f.base U.1)).op : + op U ⟶ op (opensFunctor f |>.obj ((Opens.map f.base).obj U))) := PresheafedSpace.IsOpenImmersion.app_invApp f.1 U /-- A variant of `app_inv_app` that gives an `eqToHom` instead of `homOfLe`. -/ @[reassoc] -theorem app_inv_app' (U : Opens Y) (hU : (U : Set Y) ⊆ Set.range f.1.base) : - f.1.c.app (op U) ≫ H.invApp _ ((Opens.map f.1.base).obj U) = +theorem app_inv_app' (U : Opens Y) (hU : (U : Set Y) ⊆ Set.range f.base) : + f.c.app (op U) ≫ H.invApp _ ((Opens.map f.base).obj U) = Y.presheaf.map (eqToHom <| - le_antisymm (Set.image_preimage_subset f.1.base U.1) <| - (Set.image_preimage_eq_inter_range (f := f.1.base) (t := U.1)).symm ▸ + le_antisymm (Set.image_preimage_subset f.base U.1) <| + (Set.image_preimage_eq_inter_range (f := f.base) (t := U.1)).symm ▸ Set.subset_inter_iff.mpr ⟨fun _ h => h, hU⟩).op := PresheafedSpace.IsOpenImmersion.app_invApp f.1 U @@ -1234,7 +1236,7 @@ theorem ofRestrict_invApp (X : LocallyRingedSpace) {Y : TopCat} instance stalk_iso (x : X) : IsIso (f.stalkMap x) := PresheafedSpace.IsOpenImmersion.stalk_iso f.1 x -theorem to_iso [h' : Epi f.1.base] : IsIso f := by +theorem to_iso [h' : Epi f.base] : IsIso f := by suffices IsIso (LocallyRingedSpace.forgetToSheafedSpace.map f) from isIso_of_reflects_iso _ LocallyRingedSpace.forgetToSheafedSpace exact SheafedSpace.IsOpenImmersion.to_iso f.1 diff --git a/Mathlib/Geometry/RingedSpace/PresheafedSpace.lean b/Mathlib/Geometry/RingedSpace/PresheafedSpace.lean index 43d40b52b040c..f4762b2bced5c 100644 --- a/Mathlib/Geometry/RingedSpace/PresheafedSpace.lean +++ b/Mathlib/Geometry/RingedSpace/PresheafedSpace.lean @@ -157,6 +157,9 @@ instance categoryOfPresheafedSpaces : Category (PresheafedSpace C) where variable {C} +/-- Cast `Hom X Y` as an arrow `X ⟶ Y` of presheaves. -/ +abbrev Hom.toPshHom {X Y : PresheafedSpace C} (f : Hom X Y) : X ⟶ Y := f + -- Porting note (#5229): adding an `ext` lemma. @[ext (iff := false)] theorem ext {X Y : PresheafedSpace C} (α β : X ⟶ Y) (w : α.base = β.base) diff --git a/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean b/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean index 19b9450951092..bcfd0917ee844 100644 --- a/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean +++ b/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean @@ -660,7 +660,7 @@ instance ι_isOpenImmersion (i : D.J) : IsOpenImmersion (𝖣.ι i) := by instance (i j k : D.J) : PreservesLimit (cospan (𝖣.f i j) (𝖣.f i k)) forgetToSheafedSpace := inferInstance -theorem ι_jointly_surjective (x : 𝖣.glued) : ∃ (i : D.J) (y : D.U i), (𝖣.ι i).1.base y = x := +theorem ι_jointly_surjective (x : 𝖣.glued) : ∃ (i : D.J) (y : D.U i), (𝖣.ι i).base y = x := 𝖣.ι_jointly_surjective ((LocallyRingedSpace.forgetToSheafedSpace.{u} ⋙ SheafedSpace.forget CommRingCatMax.{u, u}) ⋙ forget TopCat.{u}) x From 6eec427982c8cb527cd6e3c99d067d9e2218e719 Mon Sep 17 00:00:00 2001 From: Christian Merten Date: Wed, 16 Oct 2024 11:05:16 +0000 Subject: [PATCH 015/505] feat(RingTheory/StandardSmooth): meta properties of standard smooth ring homomorphisms (#16868) In particular we show that being standard smooth is stable under composition and base change. This contribution was created as part of the AIM workshop "Formalizing algebraic geometry" in June 2024. --- Mathlib.lean | 1 + Mathlib/RingTheory/LocalProperties/Basic.lean | 38 ++++ .../RingTheory/Localization/BaseChange.lean | 9 + .../RingTheory/RingHom/StandardSmooth.lean | 168 ++++++++++++++++++ Mathlib/RingTheory/Smooth/StandardSmooth.lean | 26 +-- 5 files changed, 217 insertions(+), 25 deletions(-) create mode 100644 Mathlib/RingTheory/RingHom/StandardSmooth.lean diff --git a/Mathlib.lean b/Mathlib.lean index ad61f598310a6..5c43f1b7e519b 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4110,6 +4110,7 @@ import Mathlib.RingTheory.RingHom.FinitePresentation import Mathlib.RingTheory.RingHom.FiniteType import Mathlib.RingTheory.RingHom.Integral import Mathlib.RingTheory.RingHom.Locally +import Mathlib.RingTheory.RingHom.StandardSmooth import Mathlib.RingTheory.RingHom.Surjective import Mathlib.RingTheory.RingHomProperties import Mathlib.RingTheory.RingInvo diff --git a/Mathlib/RingTheory/LocalProperties/Basic.lean b/Mathlib/RingTheory/LocalProperties/Basic.lean index 9c1bfac28d009..7982b328a9a8c 100644 --- a/Mathlib/RingTheory/LocalProperties/Basic.lean +++ b/Mathlib/RingTheory/LocalProperties/Basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang -/ import Mathlib.RingTheory.Localization.AtPrime +import Mathlib.RingTheory.Localization.BaseChange import Mathlib.RingTheory.Localization.Submodule import Mathlib.RingTheory.RingHomProperties @@ -261,6 +262,43 @@ lemma RingHom.OfLocalizationSpanTarget.ofIsLocalization AlgEquiv.toRingEquiv_toRingHom, Localization.coe_algEquiv_symm, IsLocalization.map_comp, RingHom.comp_id] +section + +variable (hP : RingHom.StableUnderBaseChange @P) +variable {R S Rᵣ Sᵣ : Type u} [CommRing R] [CommRing S] [CommRing Rᵣ] [CommRing Sᵣ] [Algebra R Rᵣ] + [Algebra S Sᵣ] + +include hP + +/-- Let `S` be an `R`-algebra and `Sᵣ` and `Rᵣ` be the respective localizations at a submonoid +`M` of `R`. If `P` is stable under base change and `P` holds for `algebraMap R S`, then +`P` holds for `algebraMap Rᵣ Sᵣ`. -/ +lemma RingHom.StableUnderBaseChange.of_isLocalization [Algebra R S] [Algebra R Sᵣ] [Algebra Rᵣ Sᵣ] + [IsScalarTower R S Sᵣ] [IsScalarTower R Rᵣ Sᵣ] + (M : Submonoid R) [IsLocalization M Rᵣ] [IsLocalization (Algebra.algebraMapSubmonoid S M) Sᵣ] + (h : P (algebraMap R S)) : P (algebraMap Rᵣ Sᵣ) := + letI : Algebra.IsPushout R S Rᵣ Sᵣ := Algebra.isPushout_of_isLocalization M Rᵣ S Sᵣ + hP R S Rᵣ Sᵣ h + +/-- If `P` is stable under base change and holds for `f`, then `P` holds for `f` localized +at any submonoid `M` of `R`. -/ +lemma RingHom.StableUnderBaseChange.isLocalization_map (M : Submonoid R) [IsLocalization M Rᵣ] + (f : R →+* S) [IsLocalization (M.map f) Sᵣ] (hf : P f) : + P (IsLocalization.map Sᵣ f M.le_comap_map : Rᵣ →+* Sᵣ) := by + algebraize [f, IsLocalization.map (S := Rᵣ) Sᵣ f M.le_comap_map, + (IsLocalization.map (S := Rᵣ) Sᵣ f M.le_comap_map).comp (algebraMap R Rᵣ)] + haveI : IsScalarTower R S Sᵣ := IsScalarTower.of_algebraMap_eq' + (IsLocalization.map_comp M.le_comap_map) + haveI : IsLocalization (Algebra.algebraMapSubmonoid S M) Sᵣ := + inferInstanceAs <| IsLocalization (M.map f) Sᵣ + apply hP.of_isLocalization M hf + +lemma RingHom.StableUnderBaseChange.localizationPreserves : LocalizationPreserves P := by + introv R hf + exact hP.isLocalization_map _ _ hf + +end + end RingHom end Properties diff --git a/Mathlib/RingTheory/Localization/BaseChange.lean b/Mathlib/RingTheory/Localization/BaseChange.lean index fcb23d52941ff..999592d5ff587 100644 --- a/Mathlib/RingTheory/Localization/BaseChange.lean +++ b/Mathlib/RingTheory/Localization/BaseChange.lean @@ -47,3 +47,12 @@ theorem isLocalizedModule_iff_isBaseChange : IsLocalizedModule S f ↔ IsBaseCha rw [LinearMap.coe_comp, LinearEquiv.coe_coe, Function.comp_apply, LinearEquiv.restrictScalars_apply, LinearEquiv.trans_apply, IsBaseChange.equiv_symm_apply, IsBaseChange.equiv_tmul, one_smul] + +variable (T B : Type*) [CommSemiring T] [CommSemiring B] + [Algebra R T] [Algebra T B] [Algebra R B] [Algebra A B] [IsScalarTower R T B] + [IsScalarTower R A B] + +lemma Algebra.isPushout_of_isLocalization [IsLocalization (Algebra.algebraMapSubmonoid T S) B] : + Algebra.IsPushout R T A B := by + rw [Algebra.IsPushout.comm, Algebra.isPushout_iff] + apply IsLocalizedModule.isBaseChange S diff --git a/Mathlib/RingTheory/RingHom/StandardSmooth.lean b/Mathlib/RingTheory/RingHom/StandardSmooth.lean new file mode 100644 index 0000000000000..8b47fbfe7bf3a --- /dev/null +++ b/Mathlib/RingTheory/RingHom/StandardSmooth.lean @@ -0,0 +1,168 @@ +/- +Copyright (c) 2024 Christian Merten. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Christian Merten +-/ +import Mathlib.RingTheory.LocalProperties.Basic +import Mathlib.RingTheory.Smooth.StandardSmooth +import Mathlib.Tactic.Algebraize + +/-! +# Standard smooth ring homomorphisms + +In this file we define standard smooth ring homomorphisms and show their +meta properties. + +## Notes + +This contribution was created as part of the AIM workshop "Formalizing algebraic geometry" +in June 2024. + +-/ + +universe t t' w w' u v + +variable (n m : ℕ) + +open TensorProduct + +namespace RingHom + +variable {R : Type u} {S : Type v} [CommRing R] [CommRing S] + +/-- A ring homomorphism `R →+* S` is standard smooth if `S` is standard smooth as `R`-algebra. -/ +def IsStandardSmooth (f : R →+* S) : Prop := + @Algebra.IsStandardSmooth.{t, w} _ _ _ _ f.toAlgebra + +/-- A ring homomorphism `R →+* S` is standard smooth of relative dimension `n` if +`S` is standard smooth of relative dimension `n` as `R`-algebra. -/ +def IsStandardSmoothOfRelativeDimension (f : R →+* S) : Prop := + @Algebra.IsStandardSmoothOfRelativeDimension.{t, w} n _ _ _ _ f.toAlgebra + +lemma IsStandardSmoothOfRelativeDimension.isStandardSmooth (f : R →+* S) + (hf : IsStandardSmoothOfRelativeDimension.{t, w} n f) : + IsStandardSmooth.{t, w} f := by + algebraize [f] + letI : Algebra.IsStandardSmoothOfRelativeDimension.{t, w} n R S := hf + exact Algebra.IsStandardSmoothOfRelativeDimension.isStandardSmooth n + +variable {n m} + +variable (R) in +lemma IsStandardSmoothOfRelativeDimension.id : + IsStandardSmoothOfRelativeDimension.{t, w} 0 (RingHom.id R) := + Algebra.IsStandardSmoothOfRelativeDimension.id R + +lemma IsStandardSmoothOfRelativeDimension.equiv (e : R ≃+* S) : + IsStandardSmoothOfRelativeDimension.{t, w} 0 (e : R →+* S) := by + algebraize [e.toRingHom] + exact Algebra.IsStandardSmoothOfRelativeDimension.of_algebraMap_bijective e.bijective + +variable {T : Type*} [CommRing T] + +lemma IsStandardSmooth.comp {g : S →+* T} {f : R →+* S} + (hg : IsStandardSmooth.{t', w'} g) (hf : IsStandardSmooth.{t, w} f) : + IsStandardSmooth.{max t t', max w w'} (g.comp f) := by + rw [IsStandardSmooth] + algebraize [f, g, (g.comp f)] + letI : Algebra.IsStandardSmooth R S := hf + letI : Algebra.IsStandardSmooth S T := hg + exact Algebra.IsStandardSmooth.trans.{t, t', w, w'} R S T + +lemma IsStandardSmoothOfRelativeDimension.comp {g : S →+* T} {f : R →+* S} + (hg : IsStandardSmoothOfRelativeDimension.{t', w'} n g) + (hf : IsStandardSmoothOfRelativeDimension.{t, w} m f) : + IsStandardSmoothOfRelativeDimension.{max t t', max w w'} (n + m) (g.comp f) := by + rw [IsStandardSmoothOfRelativeDimension] + algebraize [f, g, (g.comp f)] + letI : Algebra.IsStandardSmoothOfRelativeDimension m R S := hf + letI : Algebra.IsStandardSmoothOfRelativeDimension n S T := hg + exact Algebra.IsStandardSmoothOfRelativeDimension.trans m n R S T + +lemma isStandardSmooth_stableUnderComposition : + StableUnderComposition @IsStandardSmooth.{t, w} := + fun _ _ _ _ _ _ _ _ hf hg ↦ hg.comp hf + +lemma isStandardSmooth_respectsIso : RespectsIso @IsStandardSmooth.{t, w} := by + apply isStandardSmooth_stableUnderComposition.respectsIso + introv + exact (IsStandardSmoothOfRelativeDimension.equiv e).isStandardSmooth + +lemma isStandardSmoothOfRelativeDimension_respectsIso : + RespectsIso (@IsStandardSmoothOfRelativeDimension.{t, w} n) where + left {R S T _ _ _} f e hf := by + rw [← zero_add n] + exact (IsStandardSmoothOfRelativeDimension.equiv e).comp hf + right {R S T _ _ _} f e hf := by + rw [← add_zero n] + exact hf.comp (IsStandardSmoothOfRelativeDimension.equiv e) + +lemma isStandardSmooth_stableUnderBaseChange : StableUnderBaseChange @IsStandardSmooth.{t, w} := by + apply StableUnderBaseChange.mk + · exact isStandardSmooth_respectsIso + · introv h + replace h : Algebra.IsStandardSmooth R T := by + rw [RingHom.IsStandardSmooth] at h; convert h; ext; simp_rw [Algebra.smul_def]; rfl + suffices Algebra.IsStandardSmooth S (S ⊗[R] T) by + rw [RingHom.IsStandardSmooth]; convert this; ext; simp_rw [Algebra.smul_def]; rfl + infer_instance + +variable (n) + +lemma isStandardSmoothOfRelativeDimension_stableUnderBaseChange : + StableUnderBaseChange (@IsStandardSmoothOfRelativeDimension.{t, w} n) := by + apply StableUnderBaseChange.mk + · exact isStandardSmoothOfRelativeDimension_respectsIso + · introv h + replace h : Algebra.IsStandardSmoothOfRelativeDimension n R T := by + rw [RingHom.IsStandardSmoothOfRelativeDimension] at h + convert h; ext; simp_rw [Algebra.smul_def]; rfl + suffices Algebra.IsStandardSmoothOfRelativeDimension n S (S ⊗[R] T) by + rw [RingHom.IsStandardSmoothOfRelativeDimension] + convert this; ext; simp_rw [Algebra.smul_def]; rfl + infer_instance + +lemma IsStandardSmoothOfRelativeDimension.algebraMap_isLocalizationAway {Rᵣ : Type*} [CommRing Rᵣ] + [Algebra R Rᵣ] (r : R) [IsLocalization.Away r Rᵣ] : + IsStandardSmoothOfRelativeDimension.{0, 0} 0 (algebraMap R Rᵣ) := by + have : (algebraMap R Rᵣ).toAlgebra = ‹Algebra R Rᵣ› := by + ext + rw [Algebra.smul_def] + rfl + rw [IsStandardSmoothOfRelativeDimension, this] + exact Algebra.IsStandardSmoothOfRelativeDimension.localization_away r + +lemma isStandardSmooth_localizationPreserves : LocalizationPreserves IsStandardSmooth.{t, w} := + isStandardSmooth_stableUnderBaseChange.localizationPreserves + +lemma isStandardSmoothOfRelativeDimension_localizationPreserves : + LocalizationPreserves (IsStandardSmoothOfRelativeDimension.{t, w} n) := + (isStandardSmoothOfRelativeDimension_stableUnderBaseChange n).localizationPreserves + +lemma isStandardSmooth_holdsForLocalizationAway : + HoldsForLocalizationAway IsStandardSmooth.{0, 0} := by + introv R h + exact (IsStandardSmoothOfRelativeDimension.algebraMap_isLocalizationAway r).isStandardSmooth + +lemma isStandardSmoothOfRelativeDimension_holdsForLocalizationAway : + HoldsForLocalizationAway (IsStandardSmoothOfRelativeDimension.{0, 0} 0) := by + introv R h + exact IsStandardSmoothOfRelativeDimension.algebraMap_isLocalizationAway r + +lemma isStandardSmooth_stableUnderCompositionWithLocalizationAway : + StableUnderCompositionWithLocalizationAway IsStandardSmooth.{0, 0} := + isStandardSmooth_stableUnderComposition.stableUnderCompositionWithLocalizationAway + isStandardSmooth_holdsForLocalizationAway + +lemma isStandardSmoothOfRelativeDimension_stableUnderCompositionWithLocalizationAway : + StableUnderCompositionWithLocalizationAway (IsStandardSmoothOfRelativeDimension.{0, 0} n) where + left _ S T _ _ _ _ s _ _ hf := + have : (algebraMap S T).IsStandardSmoothOfRelativeDimension 0 := + IsStandardSmoothOfRelativeDimension.algebraMap_isLocalizationAway s + zero_add n ▸ IsStandardSmoothOfRelativeDimension.comp this hf + right R S _ _ _ _ _ r _ _ hf := + have : (algebraMap R S).IsStandardSmoothOfRelativeDimension 0 := + IsStandardSmoothOfRelativeDimension.algebraMap_isLocalizationAway r + add_zero n ▸ IsStandardSmoothOfRelativeDimension.comp hf this + +end RingHom diff --git a/Mathlib/RingTheory/Smooth/StandardSmooth.lean b/Mathlib/RingTheory/Smooth/StandardSmooth.lean index 988a61b9c879e..8d4159379686e 100644 --- a/Mathlib/RingTheory/Smooth/StandardSmooth.lean +++ b/Mathlib/RingTheory/Smooth/StandardSmooth.lean @@ -89,8 +89,7 @@ variable (n m : ℕ) namespace Algebra -variable (R : Type u) [CommRing R] -variable (S : Type v) [CommRing S] [Algebra R S] +variable (R : Type u) (S : Type v) [CommRing R] [CommRing S] [Algebra R S] /-- A `PreSubmersivePresentation` of an `R`-algebra `S` is a `Presentation` @@ -556,26 +555,3 @@ instance IsStandardSmoothOfRelativeDimension.baseChange end BaseChange end Algebra - -namespace RingHom - -variable {R : Type u} [CommRing R] -variable {S : Type v} [CommRing S] - -/-- A ring homomorphism `R →+* S` is standard smooth if `S` is standard smooth as `R`-algebra. -/ -def IsStandardSmooth (f : R →+* S) : Prop := - @Algebra.IsStandardSmooth.{t, w} _ _ _ _ f.toAlgebra - -/-- A ring homomorphism `R →+* S` is standard smooth of relative dimension `n` if -`S` is standard smooth of relative dimension `n` as `R`-algebra. -/ -def IsStandardSmoothOfRelativeDimension (f : R →+* S) : Prop := - @Algebra.IsStandardSmoothOfRelativeDimension.{t, w} n _ _ _ _ f.toAlgebra - -lemma IsStandardSmoothOfRelativeDimension.isStandardSmooth (f : R →+* S) - (hf : IsStandardSmoothOfRelativeDimension.{t, w} n f) : - IsStandardSmooth.{t, w} f := - letI : Algebra R S := f.toAlgebra - letI : Algebra.IsStandardSmoothOfRelativeDimension.{t, w} n R S := hf - Algebra.IsStandardSmoothOfRelativeDimension.isStandardSmooth n - -end RingHom From fd745ec84d75a14cf4e28e4ee85ef766290c778e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Wed, 16 Oct 2024 11:05:17 +0000 Subject: [PATCH 016/505] feat(SetTheory/Cardinal/Aleph): define initial ordinals (#16964) --- Mathlib/SetTheory/Cardinal/Aleph.lean | 57 ++++++++++++++++++++++- Mathlib/SetTheory/Ordinal/Arithmetic.lean | 8 ---- Mathlib/SetTheory/Ordinal/Basic.lean | 8 ++++ 3 files changed, 64 insertions(+), 9 deletions(-) diff --git a/Mathlib/SetTheory/Cardinal/Aleph.lean b/Mathlib/SetTheory/Cardinal/Aleph.lean index c1fc5cb463cb7..7a654d0363135 100644 --- a/Mathlib/SetTheory/Cardinal/Aleph.lean +++ b/Mathlib/SetTheory/Cardinal/Aleph.lean @@ -41,10 +41,65 @@ open Function Set Cardinal Equiv Order Ordinal universe u v w -namespace Cardinal +/-! ### Omega ordinals -/ + +namespace Ordinal + +/-- An ordinal is initial when it is the first ordinal with a given cardinality. + +This is written as `o.card.ord = o`, i.e. `o` is the smallest ordinal with cardinality `o.card`. -/ +def IsInitial (o : Ordinal) : Prop := + o.card.ord = o + +theorem IsInitial.ord_card {o : Ordinal} (h : IsInitial o) : o.card.ord = o := h + +theorem IsInitial.card_le_card {a b : Ordinal} (ha : IsInitial a) : a.card ≤ b.card ↔ a ≤ b := by + refine ⟨fun h ↦ ?_, Ordinal.card_le_card⟩ + rw [← ord_le_ord, ha.ord_card] at h + exact h.trans (ord_card_le b) + +theorem IsInitial.card_lt_card {a b : Ordinal} (hb : IsInitial b) : a.card < b.card ↔ a < b := + lt_iff_lt_of_le_iff_le hb.card_le_card + +theorem isInitial_ord (c : Cardinal) : IsInitial c.ord := by + rw [IsInitial, card_ord] + +theorem isInitial_natCast (n : ℕ) : IsInitial n := by + rw [IsInitial, card_nat, ord_nat] + +theorem isInitial_zero : IsInitial 0 := by + exact_mod_cast isInitial_natCast 0 + +theorem isInitial_one : IsInitial 1 := by + exact_mod_cast isInitial_natCast 1 + +theorem isInitial_omega0 : IsInitial ω := by + rw [IsInitial, card_omega0, ord_aleph0] + +theorem not_bddAbove_isInitial : ¬ BddAbove {x | IsInitial x} := by + rintro ⟨a, ha⟩ + have := ha (isInitial_ord (succ a.card)) + rw [ord_le] at this + exact (lt_succ _).not_le this + +/-- Initial ordinals are order-isomorphic to the cardinals. -/ +@[simps!] +def isInitialIso : {x // IsInitial x} ≃o Cardinal where + toFun x := x.1.card + invFun x := ⟨x.ord, isInitial_ord _⟩ + left_inv x := Subtype.ext x.2.ord_card + right_inv x := card_ord x + map_rel_iff' {a _} := a.2.card_le_card + +-- TODO: define `omega` as the enumerator function of `IsInitial`, and redefine +-- `aleph x = (omega x).card`. + +end Ordinal /-! ### Aleph cardinals -/ +namespace Cardinal + /-- The `aleph'` function gives the cardinals listed by their ordinal index. `aleph' n = n`, `aleph' ω = ℵ₀`, `aleph' (ω + 1) = succ ℵ₀`, etc. diff --git a/Mathlib/SetTheory/Ordinal/Arithmetic.lean b/Mathlib/SetTheory/Ordinal/Arithmetic.lean index aa95e0d9ce216..d8239b6b16acc 100644 --- a/Mathlib/SetTheory/Ordinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Ordinal/Arithmetic.lean @@ -2219,14 +2219,6 @@ namespace Cardinal open Ordinal -@[simp] -theorem ord_aleph0 : ord.{u} ℵ₀ = ω := - le_antisymm (ord_le.2 <| le_rfl) <| - le_of_forall_lt fun o h => by - rcases Ordinal.lt_lift_iff.1 h with ⟨o, rfl, h'⟩ - rw [lt_ord, ← lift_card, lift_lt_aleph0, ← typein_enum (· < ·) h'] - exact lt_aleph0_iff_fintype.2 ⟨Set.fintypeLTNat _⟩ - @[simp] theorem add_one_of_aleph0_le {c} (h : ℵ₀ ≤ c) : c + 1 = c := by rw [add_comm, ← card_ord c, ← card_one, ← card_add, one_add_of_omega0_le] diff --git a/Mathlib/SetTheory/Ordinal/Basic.lean b/Mathlib/SetTheory/Ordinal/Basic.lean index f6526c9ba0f28..b0d0c08af9e08 100644 --- a/Mathlib/SetTheory/Ordinal/Basic.lean +++ b/Mathlib/SetTheory/Ordinal/Basic.lean @@ -1214,6 +1214,14 @@ theorem ord_one : ord 1 = 1 := by simpa using ord_nat 1 theorem ord_ofNat (n : ℕ) [n.AtLeastTwo] : ord (no_index (OfNat.ofNat n)) = OfNat.ofNat n := ord_nat n +@[simp] +theorem ord_aleph0 : ord.{u} ℵ₀ = ω := + le_antisymm (ord_le.2 le_rfl) <| + le_of_forall_lt fun o h => by + rcases Ordinal.lt_lift_iff.1 h with ⟨o, rfl, h'⟩ + rw [lt_ord, ← lift_card, lift_lt_aleph0, ← typein_enum (· < ·) h'] + exact lt_aleph0_iff_fintype.2 ⟨Set.fintypeLTNat _⟩ + @[simp] theorem lift_ord (c) : Ordinal.lift.{u,v} (ord c) = ord (lift.{u,v} c) := by refine le_antisymm (le_of_forall_lt fun a ha => ?_) ?_ From e0e004a2406351f254842985eb68b467c53c2cdf Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Wed, 16 Oct 2024 11:05:19 +0000 Subject: [PATCH 017/505] chore(SetTheory/Ordinal/Arithmetic): fix `Ordinal.enum` type signature (#17601) For some reason, Lean was inferring [some nonsense](https://leanprover-community.github.io/mathlib4_docs/Mathlib/SetTheory/Ordinal/Basic.html#Ordinal.enum). Unfortunately, the explicit typing is required for `simp`; see [Zulip](https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/simp.20and.20subrel). --- 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 b0d0c08af9e08..2de607c6ac406 100644 --- a/Mathlib/SetTheory/Ordinal/Basic.lean +++ b/Mathlib/SetTheory/Ordinal/Basic.lean @@ -449,7 +449,7 @@ the elements of `α`. -/ -- The explicit typing is required in order for `simp` to work properly. @[simps! symm_apply_coe] def enum (r : α → α → Prop) [IsWellOrder α r] : - @RelIso (Subtype fun o => o < type r) α (Subrel (· < · ) _) r := + @RelIso { o // o < type r } α (Subrel (· < ·) { o | o < type r }) r := (typein.principalSeg r).subrelIso @[simp] From fde83d96472bb47daa76ef4a93bd6b5bae1a1fc0 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Wed, 16 Oct 2024 11:05:20 +0000 Subject: [PATCH 018/505] feat(SetTheory/Ordinal/Principal): simpler characterization of `Principal` for monotone operations (#17742) plus some drive-by spacing fixes. Subsequent PRs will use this to golf results on additive and multiplicative principal ordinals. --- Mathlib/SetTheory/Ordinal/Principal.lean | 14 +++++++++++--- 1 file changed, 11 insertions(+), 3 deletions(-) diff --git a/Mathlib/SetTheory/Ordinal/Principal.lean b/Mathlib/SetTheory/Ordinal/Principal.lean index 58596bf0f6d53..d58c0eff2dd35 100644 --- a/Mathlib/SetTheory/Ordinal/Principal.lean +++ b/Mathlib/SetTheory/Ordinal/Principal.lean @@ -11,6 +11,7 @@ import Mathlib.SetTheory.Ordinal.FixedPoint We define principal or indecomposable ordinals, and we prove the standard properties about them. ## Main definitions and results + * `Principal`: A principal or indecomposable ordinal under some binary operation. We include 0 and any other typically excluded edge cases for simplicity. * `unbounded_principal`: Principal ordinals are unbounded. @@ -20,6 +21,7 @@ We define principal or indecomposable ordinals, and we prove the standard proper multiplicative principal ordinals. ## TODO + * Prove that exponential principal ordinals are 0, 1, 2, ω, or epsilon numbers, i.e. fixed points of `fun x ↦ ω ^ x`. -/ @@ -56,6 +58,15 @@ theorem principal_iff_principal_swap {o : Ordinal} : theorem not_principal_iff {o : Ordinal} : ¬ Principal op o ↔ ∃ a < o, ∃ b < o, o ≤ op a b := by simp [Principal] +theorem principal_iff_of_monotone {o : Ordinal} + (h₁ : ∀ a, Monotone (op a)) (h₂ : ∀ a, Monotone (Function.swap op a)): + Principal op o ↔ ∀ a < o, op a a < o := by + use fun h a ha => h ha ha + intro H a b ha hb + obtain hab | hba := le_or_lt a b + · exact (h₂ b hab).trans_lt <| H b hb + · exact (h₁ a hba.le).trans_lt <| H a ha + theorem principal_zero : Principal op 0 := fun a _ h => (Ordinal.not_lt_zero a h).elim @@ -115,7 +126,6 @@ theorem unbounded_principal (op : Ordinal → Ordinal → Ordinal) : /-! #### Additive principal ordinals -/ - theorem principal_add_one : Principal (· + ·) 1 := principal_one_iff.2 <| zero_add 0 @@ -265,7 +275,6 @@ theorem mul_principal_add_is_principal_add (a : Ordinal.{u}) {b : Ordinal.{u}} ( /-! #### Multiplicative principal ordinals -/ - theorem principal_mul_one : Principal (· * ·) 1 := by rw [principal_one_iff] exact zero_mul _ @@ -423,7 +432,6 @@ theorem mul_eq_opow_log_succ {a b : Ordinal.{u}} (ha : a ≠ 0) (hb : Principal /-! #### Exponential principal ordinals -/ - theorem principal_opow_omega0 : Principal (· ^ ·) ω := fun a b ha hb => match a, b, lt_omega0.1 ha, lt_omega0.1 hb with | _, _, ⟨m, rfl⟩, ⟨n, rfl⟩ => by From 298e0e0dda8ccdfe6eeca0b8acf8615d45c4c296 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 16 Oct 2024 11:05:21 +0000 Subject: [PATCH 019/505] =?UTF-8?q?feat:=20`card=20(=CE=B9=20=E2=86=92?= =?UTF-8?q?=E2=82=80=20=CE=B1)=20=3D=20card=20=CE=B1=20^=20card=20=CE=B9`?= =?UTF-8?q?=20(#17807)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit From PFR --- Mathlib/Data/Finsupp/Fintype.lean | 18 ++++++++++-------- 1 file changed, 10 insertions(+), 8 deletions(-) diff --git a/Mathlib/Data/Finsupp/Fintype.lean b/Mathlib/Data/Finsupp/Fintype.lean index 0f5b2e08e4986..b347815decf83 100644 --- a/Mathlib/Data/Finsupp/Fintype.lean +++ b/Mathlib/Data/Finsupp/Fintype.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Anne Baanen, Alex J. Best -/ import Mathlib.Data.Finsupp.Defs -import Mathlib.Data.Fintype.Basic +import Mathlib.Data.Fintype.BigOperators /-! @@ -14,17 +14,19 @@ Some lemmas on the combination of `Finsupp`, `Fintype` and `Infinite`. -/ +variable {ι α : Type*} [DecidableEq ι] [Fintype ι] [Zero α] [Fintype α] -noncomputable instance Finsupp.fintype {ι π : Sort _} [DecidableEq ι] [Zero π] [Fintype ι] - [Fintype π] : Fintype (ι →₀ π) := +noncomputable instance Finsupp.fintype : Fintype (ι →₀ α) := Fintype.ofEquiv _ Finsupp.equivFunOnFinite.symm -instance Finsupp.infinite_of_left {ι π : Sort _} [Nontrivial π] [Zero π] [Infinite ι] : - Infinite (ι →₀ π) := - let ⟨_, hm⟩ := exists_ne (0 : π) +instance Finsupp.infinite_of_left [Nontrivial α] [Infinite ι] : Infinite (ι →₀ α) := + let ⟨_, hm⟩ := exists_ne (0 : α) Infinite.of_injective _ <| Finsupp.single_left_injective hm -instance Finsupp.infinite_of_right {ι π : Sort _} [Infinite π] [Zero π] [Nonempty ι] : - Infinite (ι →₀ π) := +instance Finsupp.infinite_of_right [Infinite α] [Nonempty ι] : Infinite (ι →₀ α) := Infinite.of_injective (fun i => Finsupp.single (Classical.arbitrary ι) i) (Finsupp.single_injective (Classical.arbitrary ι)) + +variable (ι α) in +@[simp] lemma Fintype.card_finsupp : card (ι →₀ α) = card α ^ card ι := by + simp [card_congr Finsupp.equivFunOnFinite] From ebec1b61db40c3352391654316f3644e268fee46 Mon Sep 17 00:00:00 2001 From: mathlib4-update-dependencies-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Wed, 16 Oct 2024 11:28:18 +0000 Subject: [PATCH 020/505] chore: update Mathlib dependencies 2024-10-16 (#17816) This PR updates the Mathlib dependencies. --- lake-manifest.json | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/lake-manifest.json b/lake-manifest.json index eedfc9ff53b9b..6e5aa7951c279 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,7 +5,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "b731e84cb99d738b8d9710a0ba02bf8ec8d7fd26", + "rev": "422d1a5f608fccafeddab9748e8038ef346b59bf", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", @@ -65,7 +65,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "4b61d4abc1659f15ffda5ec24fdebc229d51d066", + "rev": "7bedaed1ef024add1e171cc17706b012a9a37802", "name": "LeanSearchClient", "manifestFile": "lake-manifest.json", "inputRev": "main", From 1aaa5b99d39e4bc03db73fd848a03062b974846f Mon Sep 17 00:00:00 2001 From: Daniel Weber Date: Wed, 16 Oct 2024 12:04:57 +0000 Subject: [PATCH 021/505] feat(Data/Fin/Basic): `intCast_val_sub_eq_ite` (#17190) Co-authored-by: Daniel Weber <55664973+Command-Master@users.noreply.github.com> --- Mathlib/Data/Fin/Basic.lean | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/Mathlib/Data/Fin/Basic.lean b/Mathlib/Data/Fin/Basic.lean index 0da69a90e9ab7..7b7b730d2913b 100644 --- a/Mathlib/Data/Fin/Basic.lean +++ b/Mathlib/Data/Fin/Basic.lean @@ -433,6 +433,10 @@ theorem val_add_eq_ite {n : ℕ} (a b : Fin n) : Nat.mod_eq_of_lt (show ↑b < n from b.2)] --- Porting note: syntactically the same as the above +lemma intCast_val_sub_eq_sub_add_ite {n : ℕ} (a b : Fin n) : + ((a - b).val : ℤ) = a.val - b.val + if b ≤ a then 0 else n := by + split <;> fin_omega + section OfNatCoe @[simp] From a7cd6bdbf63a4ad953c080bcbe3104d8acf670fc Mon Sep 17 00:00:00 2001 From: Bryan Gin-ge Chen Date: Wed, 16 Oct 2024 12:56:20 +0000 Subject: [PATCH 022/505] feat: action for labeling and closing stale PRs (#17041) Per discussion in the maintainer channel. Uses a PR branch https://github.com/actions/stale/pull/1145 of [actions/stale](https://github.com/actions/stale/) since it has a feature that will allow us to only consider PRs not sitting on the [#queue](https://bit.ly/3Ymuh0U). There's still placeholder text for the comments, but this is intended to be merged with the `debug-only` flag until we have the kinks sorted out, so we could fill in that text in a later PR. --- .github/workflows/stale.yml | 25 +++++++++++++++++++++++++ 1 file changed, 25 insertions(+) create mode 100644 .github/workflows/stale.yml diff --git a/.github/workflows/stale.yml b/.github/workflows/stale.yml new file mode 100644 index 0000000000000..ea5585c38e83e --- /dev/null +++ b/.github/workflows/stale.yml @@ -0,0 +1,25 @@ +name: 'Close stale issues and PRs' +on: + schedule: + - cron: '30 1 * * *' # every day at 01:30 UTC + workflow_dispatch: + +jobs: + stale: + runs-on: ubuntu-latest + steps: + # until https://github.com/actions/stale/pull/1145 is merged + - uses: asterisk/github-actions-stale@main-only-matching-filter + with: + debug-only: 'true' # TODO: remove in follow-up PR after testing is done! + stale-pr-label: 'stale' + stale-pr-message: 'Message to comment on stale PRs.' + close-pr-label: 'closed-due-to-inactivity' + close-pr-message: 'Comment on the staled PRs while closed' + days-before-stale: 60 + days-before-close: 120 + # search string from the Zulip #queue link at https://bit.ly/4eo6brN + # "is:open is:pr -is:draft base:master sort:updated-asc status:success -label:blocked-by-other-PR -label:merge-conflict -label:awaiting-CI -label:WIP -label:awaiting-author -label:delegated -label:auto-merge-after-CI -label:ready-to-merge -label:please-adopt -label:help-wanted -label:awaiting-zulip" + # We want PRs _not_ on the queue, so we keep "is:pr -is:draft base:master" (is:open is added by the action by default) as a prefix for all queries and then negate the rest of the params in separate queries to simulate boolean OR (see https://github.com/actions/stale/pull/1145) + # except for label:auto-merge-after-CI and label:ready-to-merge which presumably will be noticed before they go stale + only-matching-filter: '[ "is:pr -is:draft base:master -status:success", "is:pr -is:draft base:master label:blocked-by-other-PR", "is:pr -is:draft base:master label:merge-conflict", "is:pr -is:draft base:master label:awaiting-CI", "is:pr -is:draft base:master label:WIP", "is:pr -is:draft base:master label:awaiting-author", "is:pr -is:draft base:master label:delegated", "is:pr -is:draft base:master label:please-adopt", "is:pr -is:draft base:master label:help-wanted", "is:pr -is:draft base:master label:awaiting-zulip" ]' From d03e389cacbdd56f187bf32e31d3cfd401fb94c7 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Wed, 16 Oct 2024 12:56:21 +0000 Subject: [PATCH 023/505] feat(SetTheory/Ordinal/Enum): more lemmas on `enumOrd` (#17668) --- .../ConditionallyCompleteLattice/Basic.lean | 12 +++++++++++ Mathlib/SetTheory/Ordinal/Enum.lean | 20 +++++++++++++++++++ 2 files changed, 32 insertions(+) diff --git a/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean b/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean index f14480f908ed0..e0b7205db6894 100644 --- a/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean +++ b/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean @@ -1106,6 +1106,18 @@ theorem cbiInf_eq_of_not_forall {p : ι → Prop} {f : Subtype p → α} (hp : ⨅ (i) (h : p i), f ⟨i, h⟩ = iInf f ⊓ sInf ∅ := cbiSup_eq_of_not_forall (α := αᵒᵈ) hp +theorem csInf_eq_bot_of_bot_mem [OrderBot α] {s : Set α} (hs : ⊥ ∈ s) : sInf s = ⊥ := + eq_bot_iff.2 <| csInf_le (OrderBot.bddBelow s) hs + +theorem ciInf_eq_bot_of_bot_mem [OrderBot α] {f : ι → α} (hs : ⊥ ∈ range f) : iInf f = ⊥ := + csInf_eq_bot_of_bot_mem hs + +theorem csSup_eq_top_of_top_mem [OrderTop α] {s : Set α} (hs : ⊤ ∈ s) : sSup s = ⊤ := + csInf_eq_bot_of_bot_mem (α := αᵒᵈ) hs + +theorem ciInf_eq_top_of_top_mem [OrderTop α] {f : ι → α} (hs : ⊤ ∈ range f) : iSup f = ⊤ := + csSup_eq_top_of_top_mem hs + open Function variable [WellFoundedLT α] diff --git a/Mathlib/SetTheory/Ordinal/Enum.lean b/Mathlib/SetTheory/Ordinal/Enum.lean index c5a4c3b7442d6..4e5d4d38a7cf6 100644 --- a/Mathlib/SetTheory/Ordinal/Enum.lean +++ b/Mathlib/SetTheory/Ordinal/Enum.lean @@ -61,6 +61,26 @@ theorem enumOrd_mem (hs : ¬ BddAbove s) (o : Ordinal) : enumOrd s o ∈ s := theorem enumOrd_strictMono (hs : ¬ BddAbove s) : StrictMono (enumOrd s) := fun a b ↦ (enumOrd_mem_aux hs b).2 a +theorem enumOrd_injective (hs : ¬ BddAbove s) : Function.Injective (enumOrd s) := + (enumOrd_strictMono hs).injective + +theorem enumOrd_inj (hs : ¬ BddAbove s) {a b : Ordinal} : enumOrd s a = enumOrd s b ↔ a = b := + (enumOrd_injective hs).eq_iff + +theorem enumOrd_le_enumOrd (hs : ¬ BddAbove s) {a b : Ordinal} : + enumOrd s a ≤ enumOrd s b ↔ a ≤ b := + (enumOrd_strictMono hs).le_iff_le + +theorem enumOrd_lt_enumOrd (hs : ¬ BddAbove s) {a b : Ordinal} : + enumOrd s a < enumOrd s b ↔ a < b := + (enumOrd_strictMono hs).lt_iff_lt + +theorem id_le_enumOrd (hs : ¬ BddAbove s) : id ≤ enumOrd s := + (enumOrd_strictMono hs).id_le + +theorem le_enumOrd_self (hs : ¬ BddAbove s) {a} : a ≤ enumOrd s a := + (enumOrd_strictMono hs).le_apply + theorem enumOrd_succ_le (hs : ¬ BddAbove s) (ha : a ∈ s) (hb : enumOrd s b < a) : enumOrd s (succ b) ≤ a := by apply enumOrd_le_of_forall_lt ha From 15a391f3007c5c377ba21af3e0a1d53502310685 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Wed, 16 Oct 2024 12:56:22 +0000 Subject: [PATCH 024/505] feat(SetTheory/Ordinal/Exponential): `lt_omega0_opow` (#17740) --- Mathlib/SetTheory/Ordinal/Exponential.lean | 20 ++++++++++++++++++++ 1 file changed, 20 insertions(+) diff --git a/Mathlib/SetTheory/Ordinal/Exponential.lean b/Mathlib/SetTheory/Ordinal/Exponential.lean index 3fad6ff39bde5..57fe4e7e7b0e4 100644 --- a/Mathlib/SetTheory/Ordinal/Exponential.lean +++ b/Mathlib/SetTheory/Ordinal/Exponential.lean @@ -486,6 +486,26 @@ theorem add_log_le_log_mul {x y : Ordinal} (b : Ordinal) (hx : x ≠ 0) (hy : y exact mul_le_mul' (opow_log_le_self b hx) (opow_log_le_self b hy) · simpa only [log_of_left_le_one hb, zero_add] using le_rfl +theorem omega0_opow_mul_nat_lt {a b : Ordinal} (h : a < b) (n : ℕ) : ω ^ a * n < ω ^ b := by + apply lt_of_lt_of_le _ (opow_le_opow_right omega0_pos (succ_le_of_lt h)) + rw [opow_succ] + exact mul_lt_mul_of_pos_left (nat_lt_omega0 n) (opow_pos a omega0_pos) + +theorem lt_omega0_opow {a b : Ordinal} (hb : b ≠ 0) : + a < ω ^ b ↔ ∃ c < b, ∃ n : ℕ, a < ω ^ c * n := by + refine ⟨fun ha ↦ ⟨_, lt_log_of_lt_opow hb ha, ?_⟩, + fun ⟨c, hc, n, hn⟩ ↦ hn.trans (omega0_opow_mul_nat_lt hc n)⟩ + obtain ⟨n, hn⟩ := lt_omega0.1 (div_opow_log_lt a one_lt_omega0) + use n.succ + rw [natCast_succ, ← hn] + exact lt_mul_succ_div a (opow_ne_zero _ omega0_ne_zero) + +theorem lt_omega0_opow_succ {a b : Ordinal} : a < ω ^ succ b ↔ ∃ n : ℕ, a < ω ^ b * n := by + refine ⟨fun ha ↦ ?_, fun ⟨n, hn⟩ ↦ hn.trans (omega0_opow_mul_nat_lt (lt_succ b) n)⟩ + obtain ⟨c, hc, n, hn⟩ := (lt_omega0_opow (succ_ne_zero b)).1 ha + refine ⟨n, hn.trans_le (mul_le_mul_right' ?_ _)⟩ + rwa [opow_le_opow_iff_right one_lt_omega0, ← lt_succ_iff] + /-! ### Interaction with `Nat.cast` -/ @[simp, norm_cast] From 8d09169a88254125170bdc9b4ffd787fbe0b3e9b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 16 Oct 2024 13:45:57 +0000 Subject: [PATCH 025/505] chore(BorelSpace): make `MeasurableSpace` arguments implicit (#17819) ... rather than instance arguments. This is motivated by non-canonical sigma-algebras showing up in the theory of Gibbs measures. From GibbsMeasure --- .../Constructions/BorelSpace/Order.lean | 18 ++--- .../Constructions/BorelSpace/Real.lean | 15 ++-- .../MeasureTheory/Measure/MeasureSpace.lean | 78 +++++++++---------- 3 files changed, 55 insertions(+), 56 deletions(-) diff --git a/Mathlib/MeasureTheory/Constructions/BorelSpace/Order.lean b/Mathlib/MeasureTheory/Constructions/BorelSpace/Order.lean index 1014d37263332..8a3ca4d20c7d5 100644 --- a/Mathlib/MeasureTheory/Constructions/BorelSpace/Order.lean +++ b/Mathlib/MeasureTheory/Constructions/BorelSpace/Order.lean @@ -94,8 +94,8 @@ end OrderTopology section Orders -variable [TopologicalSpace α] [MeasurableSpace α] [OpensMeasurableSpace α] -variable [MeasurableSpace δ] +variable [TopologicalSpace α] {mα : MeasurableSpace α} [OpensMeasurableSpace α] +variable {mδ : MeasurableSpace δ} section Preorder @@ -473,7 +473,7 @@ end LinearOrder section Lattice -variable [TopologicalSpace γ] [MeasurableSpace γ] [BorelSpace γ] +variable [TopologicalSpace γ] {mγ : MeasurableSpace γ} [BorelSpace γ] instance (priority := 100) ContinuousSup.measurableSup [Sup γ] [ContinuousSup γ] : MeasurableSup γ where @@ -499,9 +499,9 @@ end Orders section BorelSpace -variable [TopologicalSpace α] [MeasurableSpace α] [BorelSpace α] -variable [TopologicalSpace β] [MeasurableSpace β] [BorelSpace β] -variable [MeasurableSpace δ] +variable [TopologicalSpace α] {mα : MeasurableSpace α} [BorelSpace α] +variable [TopologicalSpace β] {mβ : MeasurableSpace β} [BorelSpace β] +variable {mδ : MeasurableSpace δ} section LinearOrder @@ -721,7 +721,7 @@ end LinearOrder section ConditionallyCompleteLattice @[measurability] -theorem Measurable.iSup_Prop {α} [MeasurableSpace α] [ConditionallyCompleteLattice α] +theorem Measurable.iSup_Prop {α} {mα : MeasurableSpace α} [ConditionallyCompleteLattice α] (p : Prop) {f : δ → α} (hf : Measurable f) : Measurable fun b => ⨆ _ : p, f b := by classical simp_rw [ciSup_eq_ite] @@ -730,7 +730,7 @@ theorem Measurable.iSup_Prop {α} [MeasurableSpace α] [ConditionallyCompleteLat · exact measurable_const @[measurability] -theorem Measurable.iInf_Prop {α} [MeasurableSpace α] [ConditionallyCompleteLattice α] +theorem Measurable.iInf_Prop {α} {mα : MeasurableSpace α} [ConditionallyCompleteLattice α] (p : Prop) {f : δ → α} (hf : Measurable f) : Measurable fun b => ⨅ _ : p, f b := by classical simp_rw [ciInf_eq_ite] @@ -923,7 +923,7 @@ section ENNReal /-- One can cut out `ℝ≥0∞` into the sets `{0}`, `Ico (t^n) (t^(n+1))` for `n : ℤ` and `{∞}`. This gives a way to compute the measure of a set in terms of sets on which a given function `f` does not fluctuate by more than `t`. -/ -theorem measure_eq_measure_preimage_add_measure_tsum_Ico_zpow {α : Type*} [MeasurableSpace α] +theorem measure_eq_measure_preimage_add_measure_tsum_Ico_zpow {α : Type*} {mα : MeasurableSpace α} (μ : Measure α) {f : α → ℝ≥0∞} (hf : Measurable f) {s : Set α} (hs : MeasurableSet s) {t : ℝ≥0} (ht : 1 < t) : μ s = diff --git a/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean b/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean index c5fadec87d50e..2f7ad05237166 100644 --- a/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean +++ b/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean @@ -130,7 +130,7 @@ theorem measure_ext_Ioo_rat {μ ν : Measure ℝ} [IsLocallyFiniteMeasure μ] end Real -variable [MeasurableSpace α] +variable {mα : MeasurableSpace α} @[measurability, fun_prop] theorem measurable_real_toNNReal : Measurable Real.toNNReal := @@ -217,7 +217,7 @@ def ennrealEquivSum : ℝ≥0∞ ≃ᵐ ℝ≥0 ⊕ Unit := open Function (uncurry) -theorem measurable_of_measurable_nnreal_prod [MeasurableSpace β] [MeasurableSpace γ] +theorem measurable_of_measurable_nnreal_prod {_ : MeasurableSpace β} {_ : MeasurableSpace γ} {f : ℝ≥0∞ × β → γ} (H₁ : Measurable fun p : ℝ≥0 × β => f (p.1, p.2)) (H₂ : Measurable fun x => f (∞, x)) : Measurable f := let e : ℝ≥0∞ × β ≃ᵐ (ℝ≥0 × β) ⊕ (Unit × β) := @@ -225,7 +225,7 @@ theorem measurable_of_measurable_nnreal_prod [MeasurableSpace β] [MeasurableSpa (MeasurableEquiv.sumProdDistrib _ _ _) e.symm.measurable_comp_iff.1 <| measurable_sum H₁ (H₂.comp measurable_id.snd) -theorem measurable_of_measurable_nnreal_nnreal [MeasurableSpace β] {f : ℝ≥0∞ × ℝ≥0∞ → β} +theorem measurable_of_measurable_nnreal_nnreal {_ : MeasurableSpace β} {f : ℝ≥0∞ × ℝ≥0∞ → β} (h₁ : Measurable fun p : ℝ≥0 × ℝ≥0 => f (p.1, p.2)) (h₂ : Measurable fun r : ℝ≥0 => f (∞, r)) (h₃ : Measurable fun r : ℝ≥0 => f (r, ∞)) : Measurable f := measurable_of_measurable_nnreal_prod @@ -382,8 +382,8 @@ theorem AEMeasurable.ennreal_tsum {ι} [Countable ι] {f : ι → α → ℝ≥0 exact fun s => Finset.aemeasurable_sum s fun i _ => h i @[measurability, fun_prop] -theorem AEMeasurable.nnreal_tsum {α : Type*} [MeasurableSpace α] {ι : Type*} [Countable ι] - {f : ι → α → NNReal} {μ : MeasureTheory.Measure α} (h : ∀ i : ι, AEMeasurable (f i) μ) : +theorem AEMeasurable.nnreal_tsum {α : Type*} {_ : MeasurableSpace α} {ι : Type*} [Countable ι] + {f : ι → α → NNReal} {μ : Measure α} (h : ∀ i : ι, AEMeasurable (f i) μ) : AEMeasurable (fun x : α => ∑' i : ι, f i x) μ := by simp_rw [NNReal.tsum_eq_toNNReal_tsum] exact (AEMeasurable.ennreal_tsum fun i => (h i).coe_nnreal_ennreal).ennreal_toNNReal @@ -470,9 +470,8 @@ end NNReal spanning measurable sets with finite measure on which `f` is bounded. See also `StronglyMeasurable.exists_spanning_measurableSet_norm_le` for functions into normed groups. -/ --- We redeclare `α` to temporarily avoid the `[MeasurableSpace α]` instance. -theorem exists_spanning_measurableSet_le {α : Type*} {m : MeasurableSpace α} {f : α → ℝ≥0} - (hf : Measurable f) (μ : Measure α) [SigmaFinite μ] : +theorem exists_spanning_measurableSet_le {f : α → ℝ≥0} (hf : Measurable f) (μ : Measure α) + [SigmaFinite μ] : ∃ s : ℕ → Set α, (∀ n, MeasurableSet (s n) ∧ μ (s n) < ∞ ∧ ∀ x ∈ s n, f x ≤ n) ∧ ⋃ i, s i = Set.univ := by diff --git a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean index 6068dcd58e274..7a7ab5f8c32c6 100644 --- a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean +++ b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean @@ -171,7 +171,7 @@ theorem measure_biUnion_finset {s : Finset ι} {f : ι → Set α} (hd : Pairwis /-- The measure of an a.e. disjoint union (even uncountable) of null-measurable sets is at least the sum of the measures of the sets. -/ -theorem tsum_meas_le_meas_iUnion_of_disjoint₀ {ι : Type*} [MeasurableSpace α] (μ : Measure α) +theorem tsum_meas_le_meas_iUnion_of_disjoint₀ {ι : Type*} {_ : MeasurableSpace α} (μ : Measure α) {As : ι → Set α} (As_mble : ∀ i : ι, NullMeasurableSet (As i) μ) (As_disj : Pairwise (AEDisjoint μ on As)) : (∑' i, μ (As i)) ≤ μ (⋃ i, As i) := by rw [ENNReal.tsum_eq_iSup_sum, iSup_le_iff] @@ -182,7 +182,7 @@ theorem tsum_meas_le_meas_iUnion_of_disjoint₀ {ι : Type*} [MeasurableSpace α /-- The measure of a disjoint union (even uncountable) of measurable sets is at least the sum of the measures of the sets. -/ -theorem tsum_meas_le_meas_iUnion_of_disjoint {ι : Type*} [MeasurableSpace α] (μ : Measure α) +theorem tsum_meas_le_meas_iUnion_of_disjoint {ι : Type*} {_ : MeasurableSpace α} (μ : Measure α) {As : ι → Set α} (As_mble : ∀ i : ι, MeasurableSet (As i)) (As_disj : Pairwise (Disjoint on As)) : (∑' i, μ (As i)) ≤ μ (⋃ i, As i) := tsum_meas_le_meas_iUnion_of_disjoint₀ μ (fun i ↦ (As_mble i).nullMeasurableSet) @@ -518,7 +518,7 @@ theorem measure_iInter_eq_iInf [Countable ι] {s : ι → Set α} (h : ∀ i, Nu /-- Continuity from above: the measure of the intersection of a sequence of measurable sets is the infimum of the measures of the partial intersections. -/ -theorem measure_iInter_eq_iInf' {α ι : Type*} [MeasurableSpace α] {μ : Measure α} +theorem measure_iInter_eq_iInf' {α ι : Type*} {_ : MeasurableSpace α} {μ : Measure α} [Countable ι] [Preorder ι] [IsDirected ι (· ≤ ·)] {f : ι → Set α} (h : ∀ i, NullMeasurableSet (f i) μ) (hfin : ∃ i, μ (f i) ≠ ∞) : μ (⋂ i, f i) = ⨅ i, μ (⋂ j ≤ i, f j) := by @@ -560,7 +560,7 @@ theorem tendsto_measure_iUnion_atBot [Preorder ι] [IsCountablyGenerated (atBot sets is the limit of the measures of the partial unions. -/ theorem tendsto_measure_iUnion_accumulate {α ι : Type*} [Preorder ι] [IsCountablyGenerated (atTop : Filter ι)] - [MeasurableSpace α] {μ : Measure α} {f : ι → Set α} : + {_ : MeasurableSpace α} {μ : Measure α} {f : ι → Set α} : Tendsto (fun i ↦ μ (Accumulate f i)) atTop (𝓝 (μ (⋃ i, f i))) := by refine .of_neBot_imp fun h ↦ ?_ have := (atTop_neBot_iff.1 h).2 @@ -582,7 +582,7 @@ theorem tendsto_measure_iInter [Countable ι] [Preorder ι] {s : ι → Set α} /-- Continuity from above: the measure of the intersection of a sequence of measurable sets such that one has finite measure is the limit of the measures of the partial intersections. -/ -theorem tendsto_measure_iInter' {α ι : Type*} [MeasurableSpace α] {μ : Measure α} [Countable ι] +theorem tendsto_measure_iInter' {α ι : Type*} {_ : MeasurableSpace α} {μ : Measure α} [Countable ι] [Preorder ι] [IsDirected ι (· ≤ ·)] {f : ι → Set α} (hm : ∀ i, NullMeasurableSet (f i) μ) (hf : ∃ i, μ (f i) ≠ ∞) : Tendsto (fun i ↦ μ (⋂ j ∈ {j | j ≤ i}, f j)) atTop (𝓝 (μ (⋂ i, f i))) := by @@ -715,7 +715,7 @@ theorem measure_toMeasurable_inter {s t : Set α} (hs : MeasurableSet s) (ht : /-! ### The `ℝ≥0∞`-module of measures -/ -instance instZero [MeasurableSpace α] : Zero (Measure α) := +instance instZero {_ : MeasurableSpace α} : Zero (Measure α) := ⟨{ toOuterMeasure := 0 m_iUnion := fun _f _hf _hd => tsum_zero.symm trim_le := OuterMeasure.trim_zero.le }⟩ @@ -745,10 +745,10 @@ instance instSubsingleton [IsEmpty α] {m : MeasurableSpace α} : Subsingleton ( theorem eq_zero_of_isEmpty [IsEmpty α] {_m : MeasurableSpace α} (μ : Measure α) : μ = 0 := Subsingleton.elim μ 0 -instance instInhabited [MeasurableSpace α] : Inhabited (Measure α) := +instance instInhabited {_ : MeasurableSpace α} : Inhabited (Measure α) := ⟨0⟩ -instance instAdd [MeasurableSpace α] : Add (Measure α) := +instance instAdd {_ : MeasurableSpace α} : Add (Measure α) := ⟨fun μ₁ μ₂ => { toOuterMeasure := μ₁.toOuterMeasure + μ₂.toOuterMeasure m_iUnion := fun s hs hd => @@ -774,7 +774,7 @@ section SMul variable [SMul R ℝ≥0∞] [IsScalarTower R ℝ≥0∞ ℝ≥0∞] variable [SMul R' ℝ≥0∞] [IsScalarTower R' ℝ≥0∞ ℝ≥0∞] -instance instSMul [MeasurableSpace α] : SMul R (Measure α) := +instance instSMul {_ : MeasurableSpace α} : SMul R (Measure α) := ⟨fun c μ => { toOuterMeasure := c • μ.toOuterMeasure m_iUnion := fun s hs hd => by @@ -796,15 +796,15 @@ theorem smul_apply {_m : MeasurableSpace α} (c : R) (μ : Measure α) (s : Set (c • μ) s = c • μ s := rfl -instance instSMulCommClass [SMulCommClass R R' ℝ≥0∞] [MeasurableSpace α] : +instance instSMulCommClass [SMulCommClass R R' ℝ≥0∞] {_ : MeasurableSpace α} : SMulCommClass R R' (Measure α) := ⟨fun _ _ _ => ext fun _ _ => smul_comm _ _ _⟩ -instance instIsScalarTower [SMul R R'] [IsScalarTower R R' ℝ≥0∞] [MeasurableSpace α] : +instance instIsScalarTower [SMul R R'] [IsScalarTower R R' ℝ≥0∞] {_ : MeasurableSpace α} : IsScalarTower R R' (Measure α) := ⟨fun _ _ _ => ext fun _ _ => smul_assoc _ _ _⟩ -instance instIsCentralScalar [SMul Rᵐᵒᵖ ℝ≥0∞] [IsCentralScalar R ℝ≥0∞] [MeasurableSpace α] : +instance instIsCentralScalar [SMul Rᵐᵒᵖ ℝ≥0∞] [IsCentralScalar R ℝ≥0∞] {_ : MeasurableSpace α} : IsCentralScalar R (Measure α) := ⟨fun _ _ => ext fun _ _ => op_smul_eq_smul _ _⟩ @@ -815,10 +815,10 @@ instance instNoZeroSMulDivisors [Zero R] [SMulWithZero R ℝ≥0∞] [IsScalarTo eq_zero_or_eq_zero_of_smul_eq_zero h := by simpa [Ne, ext_iff', forall_or_left] using h instance instMulAction [Monoid R] [MulAction R ℝ≥0∞] [IsScalarTower R ℝ≥0∞ ℝ≥0∞] - [MeasurableSpace α] : MulAction R (Measure α) := + {_ : MeasurableSpace α} : MulAction R (Measure α) := Injective.mulAction _ toOuterMeasure_injective smul_toOuterMeasure -instance instAddCommMonoid [MeasurableSpace α] : AddCommMonoid (Measure α) := +instance instAddCommMonoid {_ : MeasurableSpace α} : AddCommMonoid (Measure α) := toOuterMeasure_injective.addCommMonoid toOuterMeasure zero_toOuterMeasure add_toOuterMeasure fun _ _ => smul_toOuterMeasure _ _ @@ -836,12 +836,12 @@ theorem finset_sum_apply {m : MeasurableSpace α} (I : Finset ι) (μ : ι → M (∑ i ∈ I, μ i) s = ∑ i ∈ I, μ i s := by rw [coe_finset_sum, Finset.sum_apply] instance instDistribMulAction [Monoid R] [DistribMulAction R ℝ≥0∞] [IsScalarTower R ℝ≥0∞ ℝ≥0∞] - [MeasurableSpace α] : DistribMulAction R (Measure α) := + {_ : MeasurableSpace α} : DistribMulAction R (Measure α) := Injective.distribMulAction ⟨⟨toOuterMeasure, zero_toOuterMeasure⟩, add_toOuterMeasure⟩ toOuterMeasure_injective smul_toOuterMeasure -instance instModule [Semiring R] [Module R ℝ≥0∞] [IsScalarTower R ℝ≥0∞ ℝ≥0∞] [MeasurableSpace α] : - Module R (Measure α) := +instance instModule [Semiring R] [Module R ℝ≥0∞] [IsScalarTower R ℝ≥0∞ ℝ≥0∞] + {_ : MeasurableSpace α} : Module R (Measure α) := Injective.module R ⟨⟨toOuterMeasure, zero_toOuterMeasure⟩, add_toOuterMeasure⟩ toOuterMeasure_injective smul_toOuterMeasure @@ -902,7 +902,7 @@ theorem measure_toMeasurable_add_inter_right {s t : Set α} (hs : MeasurableSet /-- Measures are partially ordered. -/ -instance instPartialOrder [MeasurableSpace α] : PartialOrder (Measure α) where +instance instPartialOrder {_ : MeasurableSpace α} : PartialOrder (Measure α) where le m₁ m₂ := ∀ s, m₁ s ≤ m₂ s le_refl _ _ := le_rfl le_trans _ _ _ h₁ h₂ s := le_trans (h₁ s) (h₂ s) @@ -924,7 +924,7 @@ theorem lt_iff : μ < ν ↔ μ ≤ ν ∧ ∃ s, MeasurableSet s ∧ μ s < ν theorem lt_iff' : μ < ν ↔ μ ≤ ν ∧ ∃ s, μ s < ν s := lt_iff_le_not_le.trans <| and_congr Iff.rfl <| by simp only [le_iff', not_forall, not_le] -instance covariantAddLE [MeasurableSpace α] : +instance covariantAddLE {_ : MeasurableSpace α} : CovariantClass (Measure α) (Measure α) (· + ·) (· ≤ ·) := ⟨fun _ν _μ₁ _μ₂ hμ s => add_le_add_left (hμ s) _⟩ @@ -950,7 +950,7 @@ theorem sInf_caratheodory (s : Set α) (hs : MeasurableSet s) : rw [← measure_inter_add_diff u hs] exact add_le_add (hm <| inter_subset_inter_left _ htu) (hm <| diff_subset_diff_left htu) -instance [MeasurableSpace α] : InfSet (Measure α) := +instance {_ : MeasurableSpace α} : InfSet (Measure α) := ⟨fun m => (sInf (toOuterMeasure '' m)).toMeasure <| sInf_caratheodory⟩ theorem sInf_apply (hs : MeasurableSet s) : sInf m s = sInf (toOuterMeasure '' m) s := @@ -965,13 +965,13 @@ private theorem measure_le_sInf (h : ∀ μ' ∈ m, μ ≤ μ') : μ ≤ sInf m le_sInf <| forall_mem_image.2 fun _ hμ ↦ toOuterMeasure_le.2 <| h _ hμ le_iff.2 fun s hs => by rw [sInf_apply hs]; exact this s -instance instCompleteSemilatticeInf [MeasurableSpace α] : CompleteSemilatticeInf (Measure α) := +instance instCompleteSemilatticeInf {_ : MeasurableSpace α} : CompleteSemilatticeInf (Measure α) := { (by infer_instance : PartialOrder (Measure α)), (by infer_instance : InfSet (Measure α)) with sInf_le := fun _s _a => measure_sInf_le le_sInf := fun _s _a => measure_le_sInf } -instance instCompleteLattice [MeasurableSpace α] : CompleteLattice (Measure α) := +instance instCompleteLattice {_ : MeasurableSpace α} : CompleteLattice (Measure α) := { completeLatticeOfCompleteSemilatticeInf (Measure α) with top := { toOuterMeasure := ⊤, @@ -997,7 +997,7 @@ theorem _root_.MeasureTheory.OuterMeasure.toMeasure_top : toOuterMeasure_toMeasure (μ := ⊤) @[simp] -theorem toOuterMeasure_top [MeasurableSpace α] : +theorem toOuterMeasure_top {_ : MeasurableSpace α} : (⊤ : Measure α).toOuterMeasure = (⊤ : OuterMeasure α) := rfl @@ -1235,7 +1235,7 @@ def comapₗ [MeasurableSpace α] (f : α → β) : Measure β →ₗ[ℝ≥0∞ exact hf.2 s hs else 0 -theorem comapₗ_apply {β} [MeasurableSpace α] {mβ : MeasurableSpace β} (f : α → β) +theorem comapₗ_apply {β} {_ : MeasurableSpace α} {mβ : MeasurableSpace β} (f : α → β) (hfi : Injective f) (hf : ∀ s, MeasurableSet s → MeasurableSet (f '' s)) (μ : Measure β) (hs : MeasurableSet s) : comapₗ f μ s = μ (f '' s) := by rw [comapₗ, dif_pos, liftLinear_apply _ hs, OuterMeasure.comap_apply, coe_toOuterMeasure] @@ -1253,35 +1253,35 @@ def comap [MeasurableSpace α] (f : α → β) (μ : Measure β) : Measure α := exact (measure_inter_add_diff₀ _ (hf.2 s hs)).symm else 0 -theorem comap_apply₀ [MeasurableSpace α] (f : α → β) (μ : Measure β) (hfi : Injective f) +theorem comap_apply₀ {_ : MeasurableSpace α} (f : α → β) (μ : Measure β) (hfi : Injective f) (hf : ∀ s, MeasurableSet s → NullMeasurableSet (f '' s) μ) (hs : NullMeasurableSet s (comap f μ)) : comap f μ s = μ (f '' s) := by rw [comap, dif_pos (And.intro hfi hf)] at hs ⊢ rw [toMeasure_apply₀ _ _ hs, OuterMeasure.comap_apply, coe_toOuterMeasure] -theorem le_comap_apply {β} [MeasurableSpace α] {mβ : MeasurableSpace β} (f : α → β) (μ : Measure β) - (hfi : Injective f) (hf : ∀ s, MeasurableSet s → NullMeasurableSet (f '' s) μ) (s : Set α) : - μ (f '' s) ≤ comap f μ s := by +theorem le_comap_apply {β} {_ : MeasurableSpace α} {mβ : MeasurableSpace β} (f : α → β) + (μ : Measure β) (hfi : Injective f) (hf : ∀ s, MeasurableSet s → NullMeasurableSet (f '' s) μ) + (s : Set α) : μ (f '' s) ≤ comap f μ s := by rw [comap, dif_pos (And.intro hfi hf)] exact le_toMeasure_apply _ _ _ -theorem comap_apply {β} [MeasurableSpace α] {_mβ : MeasurableSpace β} (f : α → β) +theorem comap_apply {β} {_ : MeasurableSpace α} {_mβ : MeasurableSpace β} (f : α → β) (hfi : Injective f) (hf : ∀ s, MeasurableSet s → MeasurableSet (f '' s)) (μ : Measure β) (hs : MeasurableSet s) : comap f μ s = μ (f '' s) := comap_apply₀ f μ hfi (fun s hs => (hf s hs).nullMeasurableSet) hs.nullMeasurableSet -theorem comapₗ_eq_comap {β} [MeasurableSpace α] {_mβ : MeasurableSpace β} (f : α → β) +theorem comapₗ_eq_comap {β} {_ : MeasurableSpace α} {_mβ : MeasurableSpace β} (f : α → β) (hfi : Injective f) (hf : ∀ s, MeasurableSet s → MeasurableSet (f '' s)) (μ : Measure β) (hs : MeasurableSet s) : comapₗ f μ s = comap f μ s := (comapₗ_apply f hfi hf μ hs).trans (comap_apply f hfi hf μ hs).symm -theorem measure_image_eq_zero_of_comap_eq_zero {β} [MeasurableSpace α] {_mβ : MeasurableSpace β} +theorem measure_image_eq_zero_of_comap_eq_zero {β} {_ : MeasurableSpace α} {_mβ : MeasurableSpace β} (f : α → β) (μ : Measure β) (hfi : Injective f) (hf : ∀ s, MeasurableSet s → NullMeasurableSet (f '' s) μ) {s : Set α} (hs : comap f μ s = 0) : μ (f '' s) = 0 := le_antisymm ((le_comap_apply f μ hfi hf s).trans hs.le) (zero_le _) -theorem ae_eq_image_of_ae_eq_comap {β} [MeasurableSpace α] {mβ : MeasurableSpace β} (f : α → β) +theorem ae_eq_image_of_ae_eq_comap {β} {_ : MeasurableSpace α} {mβ : MeasurableSpace β} (f : α → β) (μ : Measure β) (hfi : Injective f) (hf : ∀ s, MeasurableSet s → NullMeasurableSet (f '' s) μ) {s t : Set α} (hst : s =ᵐ[comap f μ] t) : f '' s =ᵐ[μ] f '' t := by rw [EventuallyEq, ae_iff] at hst ⊢ @@ -1298,7 +1298,7 @@ theorem ae_eq_image_of_ae_eq_comap {β} [MeasurableSpace α] {mβ : MeasurableSp rw [h_eq_α] at hst exact measure_image_eq_zero_of_comap_eq_zero f μ hfi hf hst -theorem NullMeasurableSet.image {β} [MeasurableSpace α] {mβ : MeasurableSpace β} (f : α → β) +theorem NullMeasurableSet.image {β} {_ : MeasurableSpace α} {mβ : MeasurableSpace β} (f : α → β) (μ : Measure β) (hfi : Injective f) (hf : ∀ s, MeasurableSet s → NullMeasurableSet (f '' s) μ) {s : Set α} (hs : NullMeasurableSet s (μ.comap f)) : NullMeasurableSet (f '' s) μ := by refine ⟨toMeasurable μ (f '' toMeasurable (μ.comap f) s), measurableSet_toMeasurable _ _, ?_⟩ @@ -1309,8 +1309,8 @@ theorem NullMeasurableSet.image {β} [MeasurableSpace α] {mβ : MeasurableSpace NullMeasurableSet.toMeasurable_ae_eq hs exact ae_eq_image_of_ae_eq_comap f μ hfi hf h.symm -theorem comap_preimage {β} [MeasurableSpace α] {mβ : MeasurableSpace β} (f : α → β) (μ : Measure β) - {s : Set β} (hf : Injective f) (hf' : Measurable f) +theorem comap_preimage {β} {_ : MeasurableSpace α} {mβ : MeasurableSpace β} (f : α → β) + (μ : Measure β) {s : Set β} (hf : Injective f) (hf' : Measurable f) (h : ∀ t, MeasurableSet t → NullMeasurableSet (f '' t) μ) (hs : MeasurableSet s) : μ.comap f (f ⁻¹' s) = μ (s ∩ range f) := by rw [comap_apply₀ _ _ hf h (hf' hs).nullMeasurableSet, image_preimage_eq_inter_range] @@ -1483,7 +1483,7 @@ protected theorem refl {_m0 : MeasurableSpace α} (μ : Measure α) : μ ≪ μ protected theorem rfl : μ ≪ μ := fun _s hs => hs -instance instIsRefl [MeasurableSpace α] : IsRefl (Measure α) (· ≪ ·) := +instance instIsRefl {_ : MeasurableSpace α} : IsRefl (Measure α) (· ≪ ·) := ⟨fun _ => AbsolutelyContinuous.rfl⟩ @[simp] @@ -1706,7 +1706,7 @@ theorem exists_preimage_eq_of_preimage_ae {f : α → α} (h : QuasiMeasurePrese open Pointwise @[to_additive] -theorem smul_ae_eq_of_ae_eq {G α : Type*} [Group G] [MulAction G α] [MeasurableSpace α] +theorem smul_ae_eq_of_ae_eq {G α : Type*} [Group G] [MulAction G α] {_ : MeasurableSpace α} {s t : Set α} {μ : Measure α} (g : G) (h_qmp : QuasiMeasurePreserving (g⁻¹ • · : α → α) μ μ) (h_ae_eq : s =ᵐ[μ] t) : (g • s : Set α) =ᵐ[μ] (g • t : Set α) := by @@ -1720,7 +1720,7 @@ open Pointwise @[to_additive] theorem pairwise_aedisjoint_of_aedisjoint_forall_ne_one {G α : Type*} [Group G] [MulAction G α] - [MeasurableSpace α] {μ : Measure α} {s : Set α} + {_ : MeasurableSpace α} {μ : Measure α} {s : Set α} (h_ae_disjoint : ∀ g ≠ (1 : G), AEDisjoint μ (g • s) s) (h_qmp : ∀ g : G, QuasiMeasurePreserving (g • ·) μ μ) : Pairwise (AEDisjoint μ on fun g : G => g • s) := by @@ -1934,7 +1934,7 @@ namespace MeasurableEquiv open Equiv MeasureTheory.Measure -variable [MeasurableSpace α] [MeasurableSpace β] {μ : Measure α} {ν : Measure β} +variable {_ : MeasurableSpace α} [MeasurableSpace β] {μ : Measure α} {ν : Measure β} /-- If we map a measure along a measurable equivalence, we can compute the measure on all sets (not just the measurable ones). -/ From d3d1be67abff029258750868b89b1cfa8b22b3a4 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 16 Oct 2024 14:07:05 +0000 Subject: [PATCH 026/505] feat: post-composing independent random variables (#17811) From PFR --- Mathlib/Probability/Independence/Basic.lean | 5 +++++ Mathlib/Probability/Independence/Kernel.lean | 11 +++++++++++ 2 files changed, 16 insertions(+) diff --git a/Mathlib/Probability/Independence/Basic.lean b/Mathlib/Probability/Independence/Basic.lean index cf6f2aebad8a9..cfa5d0c19e7b7 100644 --- a/Mathlib/Probability/Independence/Basic.lean +++ b/Mathlib/Probability/Independence/Basic.lean @@ -563,6 +563,11 @@ theorem iIndepFun_iff_measure_inter_preimage_eq_mul {ι : Type*} {β : ι → Ty alias ⟨iIndepFun.measure_inter_preimage_eq_mul, _⟩ := iIndepFun_iff_measure_inter_preimage_eq_mul +nonrec lemma iIndepFun.comp {β γ : ι → Type*} {mβ : ∀ i, MeasurableSpace (β i)} + {mγ : ∀ i, MeasurableSpace (γ i)} {f : ∀ i, Ω → β i} + (h : iIndepFun mβ f μ) (g : ∀ i, β i → γ i) (hg : ∀ i, Measurable (g i)) : + iIndepFun mγ (fun i ↦ g i ∘ f i) μ := h.comp _ hg + theorem indepFun_iff_indepSet_preimage {mβ : MeasurableSpace β} {mβ' : MeasurableSpace β'} [IsZeroOrProbabilityMeasure μ] (hf : Measurable f) (hg : Measurable g) : IndepFun f g μ ↔ diff --git a/Mathlib/Probability/Independence/Kernel.lean b/Mathlib/Probability/Independence/Kernel.lean index 6986bae0a89f3..cf48a19edc339 100644 --- a/Mathlib/Probability/Independence/Kernel.lean +++ b/Mathlib/Probability/Independence/Kernel.lean @@ -907,6 +907,17 @@ theorem iIndepFun_iff_measure_inter_preimage_eq_mul {ι : Type*} {β : ι → Ty alias ⟨iIndepFun.measure_inter_preimage_eq_mul, _⟩ := iIndepFun_iff_measure_inter_preimage_eq_mul +lemma iIndepFun.comp {β γ : ι → Type*} {mβ : ∀ i, MeasurableSpace (β i)} + {mγ : ∀ i, MeasurableSpace (γ i)} {f : ∀ i, Ω → β i} + (h : iIndepFun mβ f κ μ) (g : ∀ i, β i → γ i) (hg : ∀ i, Measurable (g i)) : + iIndepFun mγ (fun i ↦ g i ∘ f i) κ μ := by + rw [iIndepFun_iff_measure_inter_preimage_eq_mul] at h ⊢ + refine fun t s hs ↦ ?_ + have := h t (sets := fun i ↦ g i ⁻¹' (s i)) (fun i a ↦ hg i (hs i a)) + filter_upwards [this] with a ha + simp_rw [Set.preimage_comp] + exact ha + theorem indepFun_iff_indepSet_preimage {mβ : MeasurableSpace β} {mβ' : MeasurableSpace β'} [IsZeroOrMarkovKernel κ] (hf : Measurable f) (hg : Measurable g) : IndepFun f g κ μ ↔ From a818532749203501896bb16bfdb4187e965cd8dd Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 16 Oct 2024 14:07:07 +0000 Subject: [PATCH 027/505] =?UTF-8?q?feat(Kernel):=20`=CE=BA.comap=20g=20=3D?= =?UTF-8?q?=20=CE=BA=20=E2=88=98=20g`=20(#17823)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit From GibbsMeasure --- Mathlib/Probability/Kernel/Composition.lean | 3 +++ 1 file changed, 3 insertions(+) diff --git a/Mathlib/Probability/Kernel/Composition.lean b/Mathlib/Probability/Kernel/Composition.lean index aa3839f4407e9..7fbfe68d509e1 100644 --- a/Mathlib/Probability/Kernel/Composition.lean +++ b/Mathlib/Probability/Kernel/Composition.lean @@ -663,6 +663,9 @@ def comap (κ : Kernel α β) (g : γ → α) (hg : Measurable g) : Kernel γ β toFun a := κ (g a) measurable' := κ.measurable.comp hg +@[simp, norm_cast] +lemma coe_comap (κ : Kernel α β) (g : γ → α) (hg : Measurable g) : κ.comap g hg = κ ∘ g := rfl + theorem comap_apply (κ : Kernel α β) (hg : Measurable g) (c : γ) : comap κ g hg c = κ (g c) := rfl From 617ce534454a10c7a537bf470ff7a2f77163e113 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 16 Oct 2024 14:29:59 +0000 Subject: [PATCH 028/505] chore(MeasureTheory): let `gcongr` know about `spanningSets` (#17824) From GibbsMeasure --- Mathlib/MeasureTheory/Measure/Typeclasses.lean | 5 +++++ 1 file changed, 5 insertions(+) diff --git a/Mathlib/MeasureTheory/Measure/Typeclasses.lean b/Mathlib/MeasureTheory/Measure/Typeclasses.lean index e84781a13f059..5dc07c2bfba72 100644 --- a/Mathlib/MeasureTheory/Measure/Typeclasses.lean +++ b/Mathlib/MeasureTheory/Measure/Typeclasses.lean @@ -668,6 +668,10 @@ def spanningSets (μ : Measure α) [SigmaFinite μ] (i : ℕ) : Set α := theorem monotone_spanningSets (μ : Measure α) [SigmaFinite μ] : Monotone (spanningSets μ) := monotone_accumulate +@[gcongr] +lemma spanningSets_mono [SigmaFinite μ] {m n : ℕ} (hmn : m ≤ n) : + spanningSets μ m ⊆ spanningSets μ n := monotone_spanningSets _ hmn + theorem measurable_spanningSets (μ : Measure α) [SigmaFinite μ] (i : ℕ) : MeasurableSet (spanningSets μ i) := MeasurableSet.iUnion fun j => MeasurableSet.iUnion fun _ => μ.toFiniteSpanningSetsIn.set_mem j @@ -676,6 +680,7 @@ theorem measure_spanningSets_lt_top (μ : Measure α) [SigmaFinite μ] (i : ℕ) μ (spanningSets μ i) < ∞ := measure_biUnion_lt_top (finite_le_nat i) fun j _ => μ.toFiniteSpanningSetsIn.finite j +@[simp] theorem iUnion_spanningSets (μ : Measure α) [SigmaFinite μ] : ⋃ i : ℕ, spanningSets μ i = univ := by simp_rw [spanningSets, iUnion_accumulate, μ.toFiniteSpanningSetsIn.spanning] From edc37a1b204745492894deeb385aaf66c0540beb Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Wed, 16 Oct 2024 14:44:33 +0000 Subject: [PATCH 029/505] chore: remove porting notes about deprecation / "not used" (#17821) These deprecated/commented-out declarations are not going to return, and if we do want to deprecate something we have a much nicer infrastructure using the `@[deprecated]` attribute. It's also been a long time since the port, so we can now just delete the remaining porting notes. Closes: #11229 --- Mathlib/Algebra/Category/Semigrp/Basic.lean | 1 - Mathlib/Algebra/GCDMonoid/Multiset.lean | 3 --- Mathlib/Analysis/InnerProductSpace/Projection.lean | 1 - Mathlib/Computability/PartrecCode.lean | 1 - Mathlib/Computability/Primrec.lean | 8 -------- Mathlib/Data/NNRat/Defs.lean | 4 ---- Mathlib/Data/Vector/Basic.lean | 6 ------ Mathlib/GroupTheory/FreeGroup/NielsenSchreier.lean | 1 - Mathlib/SetTheory/Game/Short.lean | 3 --- 9 files changed, 28 deletions(-) diff --git a/Mathlib/Algebra/Category/Semigrp/Basic.lean b/Mathlib/Algebra/Category/Semigrp/Basic.lean index a9305e7dd672d..d30f2b19f232e 100644 --- a/Mathlib/Algebra/Category/Semigrp/Basic.lean +++ b/Mathlib/Algebra/Category/Semigrp/Basic.lean @@ -44,7 +44,6 @@ namespace MagmaCat @[to_additive] instance bundledHom : BundledHom @MulHom := ⟨@MulHom.toFun, @MulHom.id, @MulHom.comp, - -- Porting note: was `@MulHom.coe_inj` which is deprecated by intros; apply @DFunLike.coe_injective, by aesop_cat, by aesop_cat⟩ -- Porting note: deriving failed for `ConcreteCategory`, diff --git a/Mathlib/Algebra/GCDMonoid/Multiset.lean b/Mathlib/Algebra/GCDMonoid/Multiset.lean index ea5a208bda10a..4bbe0597cd14b 100644 --- a/Mathlib/Algebra/GCDMonoid/Multiset.lean +++ b/Mathlib/Algebra/GCDMonoid/Multiset.lean @@ -199,9 +199,6 @@ theorem extract_gcd' (s t : Multiset α) (hs : ∃ x, x ∈ s ∧ x ≠ (0 : α) contrapose! hs exact s.gcd_eq_zero_iff.1 hs -/- Porting note: Deprecated lemmas like `map_repeat` and `eq_repeat` weren't "officially" -converted to `Multiset.replicate` format yet, so I made some ad hoc ones in `Data.Multiset.Basic` -using the originals. -/ /- Porting note: The old proof used a strange form `have := _, refine ⟨s.pmap @f (fun _ ↦ id), this, extract_gcd' s _ h this⟩,` so I rearranged the proof slightly. -/ diff --git a/Mathlib/Analysis/InnerProductSpace/Projection.lean b/Mathlib/Analysis/InnerProductSpace/Projection.lean index 28ef290eefc60..6edcc178e0c8e 100644 --- a/Mathlib/Analysis/InnerProductSpace/Projection.lean +++ b/Mathlib/Analysis/InnerProductSpace/Projection.lean @@ -621,7 +621,6 @@ section reflection variable [HasOrthogonalProjection K] --- Porting note: `bit0` is deprecated. /-- Auxiliary definition for `reflection`: the reflection as a linear equivalence. -/ def reflectionLinearEquiv : E ≃ₗ[𝕜] E := LinearEquiv.ofInvolutive diff --git a/Mathlib/Computability/PartrecCode.lean b/Mathlib/Computability/PartrecCode.lean index 77e572cb6f2a6..723eb31bfc2be 100644 --- a/Mathlib/Computability/PartrecCode.lean +++ b/Mathlib/Computability/PartrecCode.lean @@ -109,7 +109,6 @@ protected def id : Code := def curry (c : Code) (n : ℕ) : Code := comp c (pair (Code.const n) Code.id) --- Porting note: `bit0` and `bit1` are deprecated. /-- An encoding of a `Nat.Partrec.Code` as a ℕ. -/ def encodeCode : Code → ℕ | zero => 0 diff --git a/Mathlib/Computability/Primrec.lean b/Mathlib/Computability/Primrec.lean index 17cbe2ce38bec..3defa207c49ee 100644 --- a/Mathlib/Computability/Primrec.lean +++ b/Mathlib/Computability/Primrec.lean @@ -706,20 +706,12 @@ theorem nat_bodd : Primrec Nat.bodd := theorem nat_div2 : Primrec Nat.div2 := (nat_div.comp .id (const 2)).of_eq fun n => n.div2_val.symm --- Porting note: this is no longer used --- theorem nat_boddDiv2 : Primrec Nat.boddDiv2 := pair nat_bodd nat_div2 - --- Porting note: bit0 is deprecated theorem nat_double : Primrec (fun n : ℕ => 2 * n) := nat_mul.comp (const _) Primrec.id --- Porting note: bit1 is deprecated theorem nat_double_succ : Primrec (fun n : ℕ => 2 * n + 1) := nat_double |> Primrec.succ.comp --- Porting note: this is no longer used --- theorem nat_div_mod : Primrec₂ fun n k : ℕ => (n / k, n % k) := pair nat_div nat_mod - end Primrec section diff --git a/Mathlib/Data/NNRat/Defs.lean b/Mathlib/Data/NNRat/Defs.lean index 2c17c0c61fcc1..00cb00c61e506 100644 --- a/Mathlib/Data/NNRat/Defs.lean +++ b/Mathlib/Data/NNRat/Defs.lean @@ -127,8 +127,6 @@ theorem coe_mul (p q : ℚ≥0) : ((p * q : ℚ≥0) : ℚ) = p * q := @[simp] lemma num_pow (q : ℚ≥0) (n : ℕ) : (q ^ n).num = q.num ^ n := by simp [num, Int.natAbs_pow] @[simp] lemma den_pow (q : ℚ≥0) (n : ℕ) : (q ^ n).den = q.den ^ n := rfl --- Porting note: `bit0` `bit1` are deprecated, so remove these theorems. - @[simp, norm_cast] theorem coe_sub (h : q ≤ p) : ((p - q : ℚ≥0) : ℚ) = p - q := max_eq_left <| le_sub_comm.2 <| by rwa [sub_zero] @@ -287,8 +285,6 @@ theorem toNNRat_lt_iff_lt_coe {p : ℚ≥0} (hq : 0 ≤ q) : toNNRat q < p ↔ q theorem lt_toNNRat_iff_coe_lt {q : ℚ≥0} : q < toNNRat p ↔ ↑q < p := NNRat.gi.gc.lt_iff_lt --- Porting note: `bit0` `bit1` are deprecated, so remove these theorems. - theorem toNNRat_mul (hp : 0 ≤ p) : toNNRat (p * q) = toNNRat p * toNNRat q := by rcases le_total 0 q with hq | hq · ext; simp [toNNRat, hp, hq, max_eq_left, mul_nonneg] diff --git a/Mathlib/Data/Vector/Basic.lean b/Mathlib/Data/Vector/Basic.lean index 14362f975a719..67059f9ba1cbd 100644 --- a/Mathlib/Data/Vector/Basic.lean +++ b/Mathlib/Data/Vector/Basic.lean @@ -73,12 +73,6 @@ theorem mk_toList : ∀ (v : Vector α n) (h), (⟨toList v, h⟩ : Vector α n) @[simp] theorem length_val (v : Vector α n) : v.val.length = n := v.2 --- Porting note: not used in mathlib and coercions done differently in Lean 4 --- @[simp] --- theorem length_coe (v : Vector α n) : --- ((coe : { l : List α // l.length = n } → List α) v).length = n := --- v.2 - @[simp] theorem toList_map {β : Type*} (v : Vector α n) (f : α → β) : (v.map f).toList = v.toList.map f := by cases v; rfl diff --git a/Mathlib/GroupTheory/FreeGroup/NielsenSchreier.lean b/Mathlib/GroupTheory/FreeGroup/NielsenSchreier.lean index 6669bfc8d8d21..61c97a1f0113b 100644 --- a/Mathlib/GroupTheory/FreeGroup/NielsenSchreier.lean +++ b/Mathlib/GroupTheory/FreeGroup/NielsenSchreier.lean @@ -121,7 +121,6 @@ instance actionGroupoidIsFree {G A : Type u} [Group G] [IsFreeGroup G] [MulActio rcases IsFreeGroup.unique_lift f' with ⟨F', hF', uF'⟩ refine ⟨uncurry F' ?_, ?_, ?_⟩ · suffices SemidirectProduct.rightHom.comp F' = MonoidHom.id _ by - -- Porting note: `MonoidHom.ext_iff` has been deprecated. exact DFunLike.ext_iff.mp this apply IsFreeGroup.ext_hom (fun x ↦ ?_) rw [MonoidHom.comp_apply, hF'] diff --git a/Mathlib/SetTheory/Game/Short.lean b/Mathlib/SetTheory/Game/Short.lean index 03c8ecbda2c5b..37fa7839716d3 100644 --- a/Mathlib/SetTheory/Game/Short.lean +++ b/Mathlib/SetTheory/Game/Short.lean @@ -219,11 +219,8 @@ instance shortNat : ∀ n : ℕ, Short n instance shortOfNat (n : ℕ) [Nat.AtLeastTwo n] : Short (no_index (OfNat.ofNat n)) := shortNat n --- Porting note: `bit0` and `bit1` are deprecated so these instances can probably be removed. -set_option linter.deprecated false in instance shortBit0 (x : PGame.{u}) [Short x] : Short (x + x) := by infer_instance -set_option linter.deprecated false in instance shortBit1 (x : PGame.{u}) [Short x] : Short ((x + x) + 1) := shortAdd _ _ /-- Auxiliary construction of decidability instances. From 2ec6c3d15fc6dec388a4b7817b2a7cee9757ab79 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Wed, 16 Oct 2024 14:44:35 +0000 Subject: [PATCH 030/505] chore: remove some `attribute [irreducible]` that don't seem to help much (#17829) This attribute doesn't seem to impact performance, so remove it and the associated porting note. --- Mathlib/AlgebraicTopology/DoldKan/GammaCompN.lean | 9 --------- Mathlib/AlgebraicTopology/DoldKan/NCompGamma.lean | 9 --------- 2 files changed, 18 deletions(-) diff --git a/Mathlib/AlgebraicTopology/DoldKan/GammaCompN.lean b/Mathlib/AlgebraicTopology/DoldKan/GammaCompN.lean index 8b92fd134e571..b6025650b82bb 100644 --- a/Mathlib/AlgebraicTopology/DoldKan/GammaCompN.lean +++ b/Mathlib/AlgebraicTopology/DoldKan/GammaCompN.lean @@ -101,9 +101,6 @@ theorem N₁Γ₀_inv_app_f_f (K : ChainComplex C ℕ) (n : ℕ) : rw [N₁Γ₀_inv_app] apply id_comp --- Porting note (#10694): added to speed up elaboration -attribute [irreducible] N₁Γ₀ - /-- Compatibility isomorphism between `toKaroubi _ ⋙ Γ₂ ⋙ N₂` and `Γ₀ ⋙ N₁` which are functors `ChainComplex C ℕ ⥤ Karoubi (ChainComplex C ℕ)`. -/ def N₂Γ₂ToKaroubiIso : toKaroubi (ChainComplex C ℕ) ⋙ Γ₂ ⋙ N₂ ≅ Γ₀ ⋙ N₁ := @@ -142,9 +139,6 @@ lemma N₂Γ₂ToKaroubiIso_inv_app (X : ChainComplex C ℕ) : rw [Splitting.ι_desc] erw [comp_id, id_comp] --- Porting note (#10694): added to speed up elaboration -attribute [irreducible] N₂Γ₂ToKaroubiIso - /-- The counit isomorphism of the Dold-Kan equivalence for additive categories. -/ def N₂Γ₂ : Γ₂ ⋙ N₂ ≅ 𝟭 (Karoubi (ChainComplex C ℕ)) := ((whiskeringLeft _ _ _).obj (toKaroubi (ChainComplex C ℕ))).preimageIso @@ -173,9 +167,6 @@ lemma whiskerLeft_toKaroubi_N₂Γ₂_hom : dsimp only [whiskeringLeft, N₂Γ₂, Functor.preimageIso] at h ⊢ exact h --- Porting note (#10694): added to speed up elaboration -attribute [irreducible] N₂Γ₂ - theorem N₂Γ₂_compatible_with_N₁Γ₀ (K : ChainComplex C ℕ) : N₂Γ₂.hom.app ((toKaroubi _).obj K) = N₂Γ₂ToKaroubiIso.hom.app K ≫ N₁Γ₀.hom.app K := congr_app whiskerLeft_toKaroubi_N₂Γ₂_hom K diff --git a/Mathlib/AlgebraicTopology/DoldKan/NCompGamma.lean b/Mathlib/AlgebraicTopology/DoldKan/NCompGamma.lean index 1e25732cfa533..a62ceb55bcb32 100644 --- a/Mathlib/AlgebraicTopology/DoldKan/NCompGamma.lean +++ b/Mathlib/AlgebraicTopology/DoldKan/NCompGamma.lean @@ -157,9 +157,6 @@ def natTrans : (N₁ : SimplicialObject C ⥤ _) ⋙ Γ₂ ⟶ toKaroubi _ where HomologicalComplex.comp_f, AlternatingFaceMapComplex.map_f, PInfty_f_naturality_assoc, NatTrans.naturality, Splitting.IndexSet.id_fst, unop_op, len_mk] --- Porting note (#10694): added to speed up elaboration -attribute [irreducible] natTrans - end Γ₂N₁ -- Porting note: removed @[simps] attribute because it was creating timeouts @@ -177,9 +174,6 @@ lemma Γ₂N₂ToKaroubiIso_inv_app (X : SimplicialObject C) : Γ₂N₂ToKaroubiIso.inv.app X = Γ₂.map (toKaroubiCompN₂IsoN₁.inv.app X) := by simp [Γ₂N₂ToKaroubiIso] --- Porting note (#10694): added to speed up elaboration -attribute [irreducible] Γ₂N₂ToKaroubiIso - namespace Γ₂N₂ /-- The natural transformation `N₂ ⋙ Γ₂ ⟶ 𝟭 (SimplicialObject C)`. -/ @@ -194,9 +188,6 @@ theorem natTrans_app_f_app (P : Karoubi (SimplicialObject C)) : dsimp only [natTrans] simp only [whiskeringLeft_obj_preimage_app, Functor.id_map, assoc] --- Porting note (#10694): added to speed up elaboration -attribute [irreducible] natTrans - end Γ₂N₂ theorem compatibility_Γ₂N₁_Γ₂N₂_natTrans (X : SimplicialObject C) : From a54be302d6fd476755ef0d77c2e4b3a7768e9c9c Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Wed, 16 Oct 2024 14:44:36 +0000 Subject: [PATCH 031/505] chore: clean up "slow" or "speedup" porting notes (#17832) Fix and remove the porting notes that don't apply anymore (for example, if `count_heartbeats` says it's an order of magnitude or more under the heartbeat limit). --- Mathlib/Algebra/DirectSum/Ring.lean | 11 +--- Mathlib/Algebra/GCDMonoid/Basic.lean | 3 - Mathlib/Algebra/Lie/Weights/Cartan.lean | 2 - Mathlib/Combinatorics/Hindman.lean | 3 +- Mathlib/Data/Quot.lean | 10 ---- Mathlib/FieldTheory/Minpoly/Field.lean | 2 - Mathlib/GroupTheory/GroupAction/ConjAct.lean | 58 +++++--------------- Mathlib/LinearAlgebra/QuotientPi.lean | 1 - Mathlib/Topology/Homotopy/HomotopyGroup.lean | 3 +- Mathlib/Topology/Metrizable/Uniformity.lean | 1 - 10 files changed, 19 insertions(+), 75 deletions(-) diff --git a/Mathlib/Algebra/DirectSum/Ring.lean b/Mathlib/Algebra/DirectSum/Ring.lean index 39d62eca17f6e..b68930289ac05 100644 --- a/Mathlib/Algebra/DirectSum/Ring.lean +++ b/Mathlib/Algebra/DirectSum/Ring.lean @@ -222,7 +222,6 @@ private nonrec theorem one_mul (x : ⨁ i, A i) : 1 * x = x := by rw [mulHom_of_of] exact of_eq_of_gradedMonoid_eq (one_mul <| GradedMonoid.mk i xi) --- Porting note (#11083): `suffices` is very slow here. private nonrec theorem mul_one (x : ⨁ i, A i) : x * 1 = x := by suffices (mulHom A).flip One.one = AddMonoidHom.id (⨁ i, A i) from DFunLike.congr_fun this x apply addHom_ext; intro i xi @@ -230,17 +229,13 @@ private nonrec theorem mul_one (x : ⨁ i, A i) : x * 1 = x := by rw [flip_apply, mulHom_of_of] exact of_eq_of_gradedMonoid_eq (mul_one <| GradedMonoid.mk i xi) -/- Porting note: Some auxiliary statements were needed in the proof of the `suffices`, -otherwise would timeout -/ private theorem mul_assoc (a b c : ⨁ i, A i) : a * b * c = a * (b * c) := by -- (`fun a b c => a * b * c` as a bundled hom) = (`fun a b c => a * (b * c)` as a bundled hom) suffices (mulHom A).compHom.comp (mulHom A) = (AddMonoidHom.compHom flipHom <| (mulHom A).flip.compHom.comp (mulHom A)).flip by - have sol := DFunLike.congr_fun (DFunLike.congr_fun (DFunLike.congr_fun this a) b) c - have aux : ∀ a b, (mulHom A) a b = a * b := fun _ _ ↦ rfl - simp only [coe_comp, Function.comp_apply, AddMonoidHom.compHom_apply_apply, aux, flip_apply, - AddMonoidHom.flipHom_apply] at sol - exact sol + simpa only [coe_comp, Function.comp_apply, AddMonoidHom.compHom_apply_apply, flip_apply, + AddMonoidHom.flipHom_apply] + using DFunLike.congr_fun (DFunLike.congr_fun (DFunLike.congr_fun this a) b) c ext ai ax bi bx ci cx dsimp only [coe_comp, Function.comp_apply, AddMonoidHom.compHom_apply_apply, flip_apply, AddMonoidHom.flipHom_apply] diff --git a/Mathlib/Algebra/GCDMonoid/Basic.lean b/Mathlib/Algebra/GCDMonoid/Basic.lean index e0744b9bcc73a..1eb20199e470b 100644 --- a/Mathlib/Algebra/GCDMonoid/Basic.lean +++ b/Mathlib/Algebra/GCDMonoid/Basic.lean @@ -89,7 +89,6 @@ variable [CancelCommMonoidWithZero α] [NormalizationMonoid α] theorem normUnit_one : normUnit (1 : α) = 1 := normUnit_coe_units 1 --- Porting note (#11083): quite slow. Improve performance? /-- Chooses an element of each associate class, by multiplying by `normUnit` -/ def normalize : α →*₀ α where toFun x := x * normUnit x @@ -1112,7 +1111,6 @@ noncomputable def gcdMonoidOfLCM [DecidableEq α] (lcm : α → α → α) rw [mul_comm, mul_dvd_mul_iff_right h_1.2] apply ac } --- Porting note (#11083): very slow; improve performance? /-- Define `NormalizedGCDMonoid` on a structure just from the `lcm` and its properties. -/ noncomputable def normalizedGCDMonoidOfLCM [NormalizationMonoid α] [DecidableEq α] (lcm : α → α → α) (dvd_lcm_left : ∀ a b, a ∣ lcm a b) (dvd_lcm_right : ∀ a b, b ∣ lcm a b) @@ -1252,7 +1250,6 @@ namespace CommGroupWithZero variable (G₀ : Type*) [CommGroupWithZero G₀] [DecidableEq G₀] --- Porting note (#11083): very slow; improve performance? -- see Note [lower instance priority] instance (priority := 100) : NormalizedGCDMonoid G₀ where normUnit x := if h : x = 0 then 1 else (Units.mk0 x h)⁻¹ diff --git a/Mathlib/Algebra/Lie/Weights/Cartan.lean b/Mathlib/Algebra/Lie/Weights/Cartan.lean index 0ef812f460f9e..e494e0e19ba5d 100644 --- a/Mathlib/Algebra/Lie/Weights/Cartan.lean +++ b/Mathlib/Algebra/Lie/Weights/Cartan.lean @@ -105,8 +105,6 @@ def rootSpaceWeightSpaceProductAux {χ₁ χ₂ χ₃ : H → R} (hχ : χ₁ + simp only [SetLike.val_smul, smul_lie, LinearMap.coe_mk, AddHom.coe_mk, LinearMap.smul_apply, SetLike.mk_smul_mk] --- Porting note (#11083): this def is _really_ slow --- See https://github.com/leanprover-community/mathlib4/issues/5028 /-- Given a nilpotent Lie subalgebra `H ⊆ L` together with `χ₁ χ₂ : H → R`, there is a natural `R`-bilinear product of root vectors and weight vectors, compatible with the actions of `H`. -/ def rootSpaceWeightSpaceProduct (χ₁ χ₂ χ₃ : H → R) (hχ : χ₁ + χ₂ = χ₃) : diff --git a/Mathlib/Combinatorics/Hindman.lean b/Mathlib/Combinatorics/Hindman.lean index 18555bba2e080..25e45662baf0d 100644 --- a/Mathlib/Combinatorics/Hindman.lean +++ b/Mathlib/Combinatorics/Hindman.lean @@ -66,8 +66,7 @@ def Ultrafilter.semigroup {M} [Semigroup M] : Semigroup (Ultrafilter M) := { Ultrafilter.mul with mul_assoc := fun U V W => Ultrafilter.coe_inj.mp <| - -- porting note (#11083): `simp` was slow to typecheck, replaced by `simp_rw` - Filter.ext' fun p => by simp_rw [Ultrafilter.eventually_mul, mul_assoc] } + Filter.ext' fun p => by simp [Ultrafilter.eventually_mul, mul_assoc] } attribute [local instance] Ultrafilter.semigroup Ultrafilter.addSemigroup diff --git a/Mathlib/Data/Quot.lean b/Mathlib/Data/Quot.lean index a6769d8795ab8..8f524efb7730b 100644 --- a/Mathlib/Data/Quot.lean +++ b/Mathlib/Data/Quot.lean @@ -109,7 +109,6 @@ theorem liftOn_mk (a : α) (f : α → γ) (h : ∀ a₁ a₂, r a₁ a₂ → f ⟨fun hf => hf.comp Quot.exists_rep, fun hf y => let ⟨x, hx⟩ := hf y; ⟨Quot.mk _ x, hx⟩⟩ /-- Descends a function `f : α → β → γ` to quotients of `α` and `β`. -/ --- porting note (#11083): removed `@[reducible]` because it caused extremely slow `simp` protected def lift₂ (f : α → β → γ) (hr : ∀ a b₁ b₂, s b₁ b₂ → f a b₁ = f a b₂) (hs : ∀ a₁ a₂ b, r a₁ a₂ → f a₁ b = f a₂ b) (q₁ : Quot r) (q₂ : Quot s) : γ := Quot.lift (fun a ↦ Quot.lift (f a) (hr a)) @@ -122,7 +121,6 @@ theorem lift₂_mk (f : α → β → γ) (hr : ∀ a b₁ b₂, s b₁ b₂ → rfl /-- Descends a function `f : α → β → γ` to quotients of `α` and `β` and applies it. -/ --- porting note (#11083): removed `@[reducible]` because it caused extremely slow `simp` protected def liftOn₂ (p : Quot r) (q : Quot s) (f : α → β → γ) (hr : ∀ a b₁ b₂, s b₁ b₂ → f a b₁ = f a b₂) (hs : ∀ a₁ a₂ b, r a₁ a₂ → f a₁ b = f a₂ b) : γ := Quot.lift₂ f hr hs p q @@ -149,7 +147,6 @@ theorem map₂_mk (f : α → β → γ) (hr : ∀ a b₁ b₂, s b₁ b₂ → rfl /-- A binary version of `Quot.recOnSubsingleton`. -/ --- porting note (#11083): removed `@[reducible]` because it caused extremely slow `simp` @[elab_as_elim] protected def recOnSubsingleton₂ {φ : Quot r → Quot s → Sort*} [h : ∀ a b, Subsingleton (φ ⟦a⟧ ⟦b⟧)] (q₁ : Quot r) @@ -449,7 +446,6 @@ protected theorem lift_mk (f : α → β) (c) (a : α) : lift f c (mk a) = f a : rfl /-- Lift a constant function on `q : Trunc α`. -/ --- porting note (#11083): removed `@[reducible]` because it caused extremely slow `simp` protected def liftOn (q : Trunc α) (f : α → β) (c : ∀ a b : α, f a = f b) : β := lift f c q @@ -498,7 +494,6 @@ instance : LawfulMonad Trunc where variable {C : Trunc α → Sort*} /-- Recursion/induction principle for `Trunc`. -/ --- porting note (#11083): removed `@[reducible]` because it caused extremely slow `simp` @[elab_as_elim] protected def rec (f : ∀ a, C (mk a)) (h : ∀ a b : α, (Eq.ndrec (f a) (Trunc.eq (mk a) (mk b)) : C (mk b)) = f b) @@ -506,14 +501,12 @@ protected def rec (f : ∀ a, C (mk a)) Quot.rec f (fun a b _ ↦ h a b) q /-- A version of `Trunc.rec` taking `q : Trunc α` as the first argument. -/ --- porting note (#11083): removed `@[reducible]` because it caused extremely slow `simp` @[elab_as_elim] protected def recOn (q : Trunc α) (f : ∀ a, C (mk a)) (h : ∀ a b : α, (Eq.ndrec (f a) (Trunc.eq (mk a) (mk b)) : C (mk b)) = f b) : C q := Trunc.rec f h q /-- A version of `Trunc.recOn` assuming the codomain is a `Subsingleton`. -/ --- porting note (#11083)s: removed `@[reducible]` because it caused extremely slow `simp` @[elab_as_elim] protected def recOnSubsingleton [∀ a, Subsingleton (C (mk a))] (q : Trunc α) (f : ∀ a, C (mk a)) : C q := @@ -557,7 +550,6 @@ theorem surjective_Quotient_mk'' : Function.Surjective (Quotient.mk'' : α → Q /-- A version of `Quotient.liftOn` taking `{s : Setoid α}` as an implicit argument instead of an instance argument. -/ --- porting note (#11083): removed `@[reducible]` because it caused extremely slow `simp` protected def liftOn' (q : Quotient s₁) (f : α → φ) (h : ∀ a b, s₁ a b → f a = f b) : φ := Quotient.liftOn q f h @@ -573,7 +565,6 @@ protected theorem liftOn'_mk'' (f : α → φ) (h) (x : α) : /-- A version of `Quotient.liftOn₂` taking `{s₁ : Setoid α} {s₂ : Setoid β}` as implicit arguments instead of instance arguments. -/ --- porting note (#11083): removed `@[reducible]` because it caused extremely slow `simp` protected def liftOn₂' (q₁ : Quotient s₁) (q₂ : Quotient s₂) (f : α → β → γ) (h : ∀ a₁ a₂ b₁ b₂, s₁ a₁ b₁ → s₂ a₂ b₂ → f a₁ a₂ = f b₁ b₂) : γ := Quotient.liftOn₂ q₁ q₂ f h @@ -632,7 +623,6 @@ protected def recOnSubsingleton' {φ : Quotient s₁ → Sort*} [∀ a, Subsingl /-- A version of `Quotient.recOnSubsingleton₂` taking `{s₁ : Setoid α} {s₂ : Setoid α}` as implicit arguments instead of instance arguments. -/ --- porting note (#11083): removed `@[reducible]` because it caused extremely slow `simp` @[elab_as_elim] protected def recOnSubsingleton₂' {φ : Quotient s₁ → Quotient s₂ → Sort*} [∀ a b, Subsingleton (φ ⟦a⟧ ⟦b⟧)] diff --git a/Mathlib/FieldTheory/Minpoly/Field.lean b/Mathlib/FieldTheory/Minpoly/Field.lean index b129b8a0ecbc5..ad3c581e77e1e 100644 --- a/Mathlib/FieldTheory/Minpoly/Field.lean +++ b/Mathlib/FieldTheory/Minpoly/Field.lean @@ -179,8 +179,6 @@ noncomputable def Fintype.subtypeProd {E : Type*} {X : Set E} (hX : X.Finite) {L variable (F E K : Type*) [Field F] [Ring E] [CommRing K] [IsDomain K] [Algebra F E] [Algebra F K] [FiniteDimensional F E] --- Porting note (#11083): removed `noncomputable!` since it seems not to be slow in lean 4, --- though it isn't very computable in practice (since neither `finrank` nor `finBasis` are). /-- Function from Hom_K(E,L) to pi type Π (x : basis), roots of min poly of x -/ def rootsOfMinPolyPiType (φ : E →ₐ[F] K) (x : range (Module.finBasis F E : _ → E)) : diff --git a/Mathlib/GroupTheory/GroupAction/ConjAct.lean b/Mathlib/GroupTheory/GroupAction/ConjAct.lean index c1d499828409f..1f5c561c568ef 100644 --- a/Mathlib/GroupTheory/GroupAction/ConjAct.lean +++ b/Mathlib/GroupTheory/GroupAction/ConjAct.lean @@ -146,18 +146,11 @@ instance unitsScalar : SMul (ConjAct Mˣ) M where smul g h := ofConjAct g * h * theorem units_smul_def (g : ConjAct Mˣ) (h : M) : g • h = ofConjAct g * h * ↑(ofConjAct g)⁻¹ := rfl --- porting note (#11083): very slow without `simp only` and need to separate `units_smul_def` --- so that things trigger appropriately instance unitsMulDistribMulAction : MulDistribMulAction (ConjAct Mˣ) M where - one_smul := by simp only [units_smul_def, ofConjAct_one, Units.val_one, one_mul, inv_one, - mul_one, forall_const] - mul_smul := by - simp only [units_smul_def] - simp only [map_mul, Units.val_mul, mul_assoc, mul_inv_rev, forall_const, «forall»] - smul_mul := by - simp only [units_smul_def] - simp only [mul_assoc, Units.inv_mul_cancel_left, forall_const, «forall»] - smul_one := by simp [units_smul_def, mul_one, Units.mul_inv, «forall», forall_const] + one_smul := by simp [units_smul_def] + mul_smul := by simp [units_smul_def, mul_assoc] + smul_mul := by simp [units_smul_def, mul_assoc] + smul_one := by simp [units_smul_def] instance unitsSMulCommClass [SMul α M] [SMulCommClass α M M] [IsScalarTower α M M] : @@ -175,15 +168,10 @@ section Semiring variable [Semiring R] --- porting note (#11083): very slow without `simp only` and need to separate `units_smul_def` --- so that things trigger appropriately instance unitsMulSemiringAction : MulSemiringAction (ConjAct Rˣ) R := { ConjAct.unitsMulDistribMulAction with - smul_zero := by - simp only [units_smul_def, mul_zero, zero_mul, «forall», forall_const] - smul_add := by - simp only [units_smul_def] - simp only [mul_add, add_mul, forall_const, «forall»] } + smul_zero := by simp [units_smul_def] + smul_add := by simp [units_smul_def, mul_add, add_mul] } end Semiring @@ -201,15 +189,9 @@ theorem ofConjAct_zero : ofConjAct (0 : ConjAct G₀) = 0 := theorem toConjAct_zero : toConjAct (0 : G₀) = 0 := rfl --- porting note (#11083): very slow without `simp only` and need to separate `smul_def` --- so that things trigger appropriately instance mulAction₀ : MulAction (ConjAct G₀) G₀ where - one_smul := by - simp only [smul_def] - simp only [map_one, one_mul, inv_one, mul_one, forall_const] - mul_smul := by - simp only [smul_def] - simp only [map_mul, mul_assoc, mul_inv_rev, forall_const, «forall»] + one_smul := by simp [smul_def] + mul_smul := by simp [smul_def, mul_assoc] instance smulCommClass₀ [SMul α G₀] [SMulCommClass α G₀ G₀] [IsScalarTower α G₀ G₀] : SMulCommClass α (ConjAct G₀) G₀ where @@ -226,16 +208,10 @@ section DivisionRing variable [DivisionRing K] --- porting note (#11083): very slow without `simp only` and need to separate `smul_def` --- so that things trigger appropriately instance distribMulAction₀ : DistribMulAction (ConjAct K) K := { ConjAct.mulAction₀ with - smul_zero := by - simp only [smul_def] - simp only [mul_zero, zero_mul, «forall», forall_const] - smul_add := by - simp only [smul_def] - simp only [mul_add, add_mul, forall_const, «forall»] } + smul_zero := by simp [smul_def] + smul_add := by simp [smul_def, mul_add, add_mul] } end DivisionRing @@ -243,17 +219,11 @@ variable [Group G] -- todo: this file is not in good order; I will refactor this after the PR --- porting note (#11083): very slow without `simp only` and need to separate `smul_def` --- so that things trigger appropriately instance : MulDistribMulAction (ConjAct G) G where - smul_mul := by - simp only [smul_def] - simp only [mul_assoc, inv_mul_cancel_left, forall_const, «forall»] - smul_one := by simp only [smul_def, mul_one, mul_inv_cancel, «forall», forall_const] - one_smul := by simp only [smul_def, ofConjAct_one, one_mul, inv_one, mul_one, forall_const] - mul_smul := by - simp only [smul_def] - simp only [map_mul, mul_assoc, mul_inv_rev, forall_const, «forall»] + smul_mul := by simp [smul_def] + smul_one := by simp [smul_def] + one_smul := by simp [smul_def] + mul_smul := by simp [smul_def, mul_assoc] instance smulCommClass [SMul α G] [SMulCommClass α G G] [IsScalarTower α G G] : SMulCommClass α (ConjAct G) G where diff --git a/Mathlib/LinearAlgebra/QuotientPi.lean b/Mathlib/LinearAlgebra/QuotientPi.lean index e4771d9be3c91..103b03974aa55 100644 --- a/Mathlib/LinearAlgebra/QuotientPi.lean +++ b/Mathlib/LinearAlgebra/QuotientPi.lean @@ -69,7 +69,6 @@ theorem quotientPiLift_mk (p : ∀ i, Submodule R (Ms i)) (f : ∀ i, Ms i → quotientPiLift p f hf (Quotient.mk x) = fun i => f i (x i) := rfl --- Porting note (#11083): split up the definition to avoid timeouts. Still slow. namespace quotientPi_aux variable (p : ∀ i, Submodule R (Ms i)) diff --git a/Mathlib/Topology/Homotopy/HomotopyGroup.lean b/Mathlib/Topology/Homotopy/HomotopyGroup.lean index daaa41db02500..fac80ea72c606 100644 --- a/Mathlib/Topology/Homotopy/HomotopyGroup.lean +++ b/Mathlib/Topology/Homotopy/HomotopyGroup.lean @@ -248,7 +248,7 @@ def homotopyTo (i : N) {p q : Ω^ N X x} (H : p.1.HomotopyRel q.1 (Cube.boundary ((⟨_, ContinuousMap.continuous_curry⟩ : C(_, _)).comp <| (cCompInsert i).comp H.toContinuousMap.curry).uncurry --- porting note (#11083): `@[simps]` no longer too slow in Lean 4 but does not generate this lemma. +-- porting note: `@[simps]` generates this lemma but it's named `homotopyTo_apply_apply` instead theorem homotopyTo_apply (i : N) {p q : Ω^ N X x} (H : p.1.HomotopyRel q.1 <| Cube.boundary N) (t : I × I) (tₙ : I^{ j // j ≠ i }) : homotopyTo i H t tₙ = H (t.fst, Cube.insertAt i (t.snd, tₙ)) := @@ -276,7 +276,6 @@ theorem homotopicTo (i : N) {p q : Ω^ N X x} : (ContinuousMap.comp ⟨_, ContinuousMap.continuous_uncurry⟩ (ContinuousMap.comp ⟨Subtype.val, by continuity⟩ H.toContinuousMap).curry).uncurry.comp <| (ContinuousMap.id I).prodMap (Cube.splitAt i).toContinuousMap --- porting note (#11083): @[simps!] no longer too slow in Lean 4. theorem homotopicFrom (i : N) {p q : Ω^ N X x} : (toLoop i p).Homotopic (toLoop i q) → Homotopic p q := by diff --git a/Mathlib/Topology/Metrizable/Uniformity.lean b/Mathlib/Topology/Metrizable/Uniformity.lean index 3aaa2860dcb66..c4bd746dd6f0d 100644 --- a/Mathlib/Topology/Metrizable/Uniformity.lean +++ b/Mathlib/Topology/Metrizable/Uniformity.lean @@ -173,7 +173,6 @@ theorem le_two_mul_dist_ofPreNNDist (d : X → X → ℝ≥0) (dist_self : ∀ x end PseudoMetricSpace --- Porting note (#11083): this is slower than in Lean3 for some reason... /-- If `X` is a uniform space with countably generated uniformity filter, there exists a `PseudoMetricSpace` structure compatible with the `UniformSpace` structure. Use `UniformSpace.pseudoMetricSpace` or `UniformSpace.metricSpace` instead. -/ From 0b07d2aeefb43ce191fcf685dbdc36baf22e714a Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Wed, 16 Oct 2024 15:23:29 +0000 Subject: [PATCH 032/505] chore: classify porting notes about `ext` (#17809) I checked all porting notes mentioning `ext` and either fixed them (see #17810), or classified them as * #5229 if it's a new `@[ext]` lemma * #11041 if it's a regression in `ext` * #11182 if `@[ext]` had to be removed This PR should contain only modifications to comments. --- .../Category/ModuleCat/Adjunctions.lean | 2 +- .../Algebra/Homology/HomologicalComplex.lean | 2 +- Mathlib/Algebra/Lie/DirectSum.lean | 24 +++++++++---------- Mathlib/Algebra/MonoidAlgebra/Basic.lean | 10 ++++---- Mathlib/Algebra/MonoidAlgebra/Defs.lean | 12 +++++----- Mathlib/Algebra/MonoidAlgebra/Division.lean | 14 +++++------ .../Algebra/MonoidAlgebra/ToDirectSum.lean | 2 +- Mathlib/Algebra/MvPolynomial/Basic.lean | 4 ++-- Mathlib/Algebra/Polynomial/Laurent.lean | 2 +- Mathlib/AlgebraicGeometry/Restrict.lean | 4 ++-- Mathlib/AlgebraicGeometry/Spec.lean | 2 +- .../AlgebraicTopology/SimplexCategory.lean | 2 +- .../SplitSimplicialObject.lean | 2 +- .../Abelian/Pseudoelements.lean | 2 +- .../NaturalTransformation/Oplax.lean | 2 +- Mathlib/CategoryTheory/Closed/Functor.lean | 4 ++-- Mathlib/CategoryTheory/Comma/Basic.lean | 2 +- Mathlib/CategoryTheory/Elements.lean | 2 +- Mathlib/CategoryTheory/Endomorphism.lean | 2 +- Mathlib/CategoryTheory/Functor/Category.lean | 2 +- .../CategoryTheory/Idempotents/Karoubi.lean | 2 +- Mathlib/CategoryTheory/Iso.lean | 8 +++---- Mathlib/CategoryTheory/Limits/Cones.lean | 4 ++-- .../Limits/Shapes/Biproducts.lean | 4 ++-- .../Limits/Shapes/Multiequalizer.lean | 2 +- Mathlib/CategoryTheory/Monoidal/Bimod.lean | 2 +- .../Monoidal/Braided/Basic.lean | 4 ++-- Mathlib/CategoryTheory/Monoidal/Center.lean | 6 ++--- .../Monoidal/Internal/Module.lean | 10 ++++---- Mathlib/CategoryTheory/Monoidal/Mon_.lean | 2 +- .../Monoidal/NaturalTransformation.lean | 2 +- Mathlib/CategoryTheory/Pi/Basic.lean | 2 +- Mathlib/CategoryTheory/Sites/Pretopology.lean | 2 +- Mathlib/CategoryTheory/Sites/Sheaf.lean | 2 +- Mathlib/Data/PFunctor/Univariate/M.lean | 2 +- .../Geometry/RingedSpace/OpenImmersion.lean | 8 +++---- .../PresheafedSpace/HasColimits.lean | 2 +- .../Category/OmegaCompletePartialOrder.lean | 6 ++--- .../RepresentationTheory/Action/Basic.lean | 2 +- .../DedekindDomain/FiniteAdeleRing.lean | 4 ++-- Mathlib/Topology/Category/CompHaus/Basic.lean | 6 ++--- .../Topology/Category/Profinite/AsLimit.lean | 3 ++- .../Category/TopCat/Limits/Products.lean | 2 +- .../Topology/Sheaves/PresheafOfFunctions.lean | 2 +- 44 files changed, 94 insertions(+), 93 deletions(-) diff --git a/Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean b/Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean index c45921b738acd..2abebe30399c0 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean @@ -369,7 +369,7 @@ def embeddingLiftIso (F : C ⥤ D) : embedding R C ⋙ lift R F ≅ F := /-- Two `R`-linear functors out of the `R`-linear completion are isomorphic iff their compositions with the embedding functor are isomorphic. -/ --- Porting note: used to be @[ext] +-- Porting note (#11182): used to be @[ext] def ext {F G : Free R C ⥤ D} [F.Additive] [F.Linear R] [G.Additive] [G.Linear R] (α : embedding R C ⋙ F ≅ embedding R C ⋙ G) : F ≅ G := NatIso.ofComponents (fun X => α.app X) diff --git a/Mathlib/Algebra/Homology/HomologicalComplex.lean b/Mathlib/Algebra/Homology/HomologicalComplex.lean index 0b0c05b2019ed..6ea13712672bd 100644 --- a/Mathlib/Algebra/Homology/HomologicalComplex.lean +++ b/Mathlib/Algebra/Homology/HomologicalComplex.lean @@ -244,7 +244,7 @@ instance : Category (HomologicalComplex V c) where end --- Porting note: added because `Hom.ext` is not triggered automatically +-- Porting note (#5229): added because `Hom.ext` is not triggered automatically @[ext] lemma hom_ext {C D : HomologicalComplex V c} (f g : C ⟶ D) (h : ∀ i, f.f i = g.f i) : f = g := by diff --git a/Mathlib/Algebra/Lie/DirectSum.lean b/Mathlib/Algebra/Lie/DirectSum.lean index d6222f0c9db3e..830fec1d310b7 100644 --- a/Mathlib/Algebra/Lie/DirectSum.lean +++ b/Mathlib/Algebra/Lie/DirectSum.lean @@ -43,13 +43,13 @@ variable [∀ i, LieRingModule L (M i)] [∀ i, LieModule R L (M i)] instance : LieRingModule L (⨁ i, M i) where bracket x m := m.mapRange (fun _ m' => ⁅x, m'⁆) fun _ => lie_zero x add_lie x y m := by - refine DFinsupp.ext fun _ => ?_ -- Porting note: Originally `ext` + refine DFinsupp.ext fun _ => ?_ -- Porting note (#11041): Originally `ext` simp only [mapRange_apply, add_apply, add_lie] lie_add x m n := by - refine DFinsupp.ext fun _ => ?_ -- Porting note: Originally `ext` + refine DFinsupp.ext fun _ => ?_ -- Porting note (#11041): Originally `ext` simp only [mapRange_apply, add_apply, lie_add] leibniz_lie x y m := by - refine DFinsupp.ext fun _ => ?_ -- Porting note: Originally `ext` + refine DFinsupp.ext fun _ => ?_ -- Porting note (#11041): Originally `ext` simp only [mapRange_apply, lie_lie, add_apply, sub_add_cancel] @[simp] @@ -58,10 +58,10 @@ theorem lie_module_bracket_apply (x : L) (m : ⨁ i, M i) (i : ι) : ⁅x, m⁆ instance : LieModule R L (⨁ i, M i) where smul_lie t x m := by - refine DFinsupp.ext fun _ => ?_ -- Porting note: Originally `ext i` + refine DFinsupp.ext fun _ => ?_ -- Porting note (#11041): Originally `ext i` simp only [smul_lie, lie_module_bracket_apply, smul_apply] lie_smul t x m := by - refine DFinsupp.ext fun _ => ?_ -- Porting note: Originally `ext i` + refine DFinsupp.ext fun _ => ?_ -- Porting note (#11041): Originally `ext i` simp only [lie_smul, lie_module_bracket_apply, smul_apply] variable (R ι L M) @@ -70,7 +70,7 @@ variable (R ι L M) def lieModuleOf [DecidableEq ι] (j : ι) : M j →ₗ⁅R,L⁆ ⨁ i, M i := { lof R ι M j with map_lie' := fun {x m} => by - refine DFinsupp.ext fun i => ?_ -- Porting note: Originally `ext i` + refine DFinsupp.ext fun i => ?_ -- Porting note (#11041): Originally `ext i` by_cases h : j = i · rw [← h]; simp · -- This used to be the end of the proof before leanprover/lean4#2644 @@ -98,16 +98,16 @@ instance lieRing : LieRing (⨁ i, L i) := { (inferInstance : AddCommGroup _) with bracket := zipWith (fun _ => fun x y => ⁅x, y⁆) fun _ => lie_zero 0 add_lie := fun x y z => by - refine DFinsupp.ext fun _ => ?_ -- Porting note: Originally `ext` + refine DFinsupp.ext fun _ => ?_ -- Porting note (#11041): Originally `ext` simp only [zipWith_apply, add_apply, add_lie] lie_add := fun x y z => by - refine DFinsupp.ext fun _ => ?_ -- Porting note: Originally `ext` + refine DFinsupp.ext fun _ => ?_ -- Porting note (#11041): Originally `ext` simp only [zipWith_apply, add_apply, lie_add] lie_self := fun x => by - refine DFinsupp.ext fun _ => ?_ -- Porting note: Originally `ext` + refine DFinsupp.ext fun _ => ?_ -- Porting note (#11041): Originally `ext` simp only [zipWith_apply, add_apply, lie_self, zero_apply] leibniz_lie := fun x y z => by - refine DFinsupp.ext fun _ => ?_ -- Porting note: Originally `ext` + refine DFinsupp.ext fun _ => ?_ -- Porting note (#11041): Originally `ext` simp only [sub_apply, zipWith_apply, add_apply, zero_apply] apply leibniz_lie } @@ -137,7 +137,7 @@ theorem lie_of [DecidableEq ι] {i j : ι} (x : L i) (y : L j) : instance lieAlgebra : LieAlgebra R (⨁ i, L i) := { (inferInstance : Module R _) with lie_smul := fun c x y => by - refine DFinsupp.ext fun _ => ?_ -- Porting note: Originally `ext` + refine DFinsupp.ext fun _ => ?_ -- Porting note (#11041): Originally `ext` simp only [zipWith_apply, smul_apply, bracket_apply, lie_smul] } variable (R ι) @@ -148,7 +148,7 @@ def lieAlgebraOf [DecidableEq ι] (j : ι) : L j →ₗ⁅R⁆ ⨁ i, L i := { lof R ι L j with toFun := of L j map_lie' := fun {x y} => by - refine DFinsupp.ext fun i => ?_ -- Porting note: Originally `ext i` + refine DFinsupp.ext fun i => ?_ -- Porting note (#11041): Originally `ext i` by_cases h : j = i · rw [← h] -- This used to be the end of the proof before leanprover/lean4#2644 diff --git a/Mathlib/Algebra/MonoidAlgebra/Basic.lean b/Mathlib/Algebra/MonoidAlgebra/Basic.lean index c859345b543ca..9eb40ad475be0 100644 --- a/Mathlib/Algebra/MonoidAlgebra/Basic.lean +++ b/Mathlib/Algebra/MonoidAlgebra/Basic.lean @@ -90,7 +90,7 @@ def liftMagma [Module k A] [IsScalarTower k A A] [SMulCommClass k A A] : sum_single_index, Function.comp_apply, one_smul, zero_smul, MulHom.coe_comp, NonUnitalAlgHom.coe_to_mulHom] right_inv F := by - -- Porting note: `ext` → `refine nonUnitalAlgHom_ext' k (MulHom.ext fun m => ?_)` + -- Porting note (#11041): `ext` → `refine nonUnitalAlgHom_ext' k (MulHom.ext fun m => ?_)` refine nonUnitalAlgHom_ext' k (MulHom.ext fun m => ?_) simp only [NonUnitalAlgHom.coe_mk, ofMagma_apply, NonUnitalAlgHom.toMulHom_eq_coe, sum_single_index, Function.comp_apply, one_smul, zero_smul, MulHom.coe_comp, @@ -109,7 +109,7 @@ In particular this provides the instance `Algebra k (MonoidAlgebra k G)`. instance algebra {A : Type*} [CommSemiring k] [Semiring A] [Algebra k A] [Monoid G] : Algebra k (MonoidAlgebra A G) := { singleOneRingHom.comp (algebraMap k A) with - -- Porting note: `ext` → `refine Finsupp.ext fun _ => ?_` + -- Porting note (#11041): `ext` → `refine Finsupp.ext fun _ => ?_` smul_def' := fun r a => by refine Finsupp.ext fun _ => ?_ -- Porting note: Newly required. @@ -125,7 +125,7 @@ def singleOneAlgHom {A : Type*} [CommSemiring k] [Semiring A] [Algebra k A] [Mon A →ₐ[k] MonoidAlgebra A G := { singleOneRingHom with commutes' := fun r => by - -- Porting note: `ext` → `refine Finsupp.ext fun _ => ?_` + -- Porting note (#11041): `ext` → `refine Finsupp.ext fun _ => ?_` refine Finsupp.ext fun _ => ?_ simp rfl } @@ -390,7 +390,7 @@ In particular this provides the instance `Algebra k k[G]`. instance algebra [CommSemiring R] [Semiring k] [Algebra R k] [AddMonoid G] : Algebra R k[G] := { singleZeroRingHom.comp (algebraMap R k) with - -- Porting note: `ext` → `refine Finsupp.ext fun _ => ?_` + -- Porting note (#11041): `ext` → `refine Finsupp.ext fun _ => ?_` smul_def' := fun r a => by refine Finsupp.ext fun _ => ?_ -- Porting note: Newly required. @@ -405,7 +405,7 @@ instance algebra [CommSemiring R] [Semiring k] [Algebra R k] [AddMonoid G] : def singleZeroAlgHom [CommSemiring R] [Semiring k] [Algebra R k] [AddMonoid G] : k →ₐ[R] k[G] := { singleZeroRingHom with commutes' := fun r => by - -- Porting note: `ext` → `refine Finsupp.ext fun _ => ?_` + -- Porting note (#11041): `ext` → `refine Finsupp.ext fun _ => ?_` refine Finsupp.ext fun _ => ?_ simp rfl } diff --git a/Mathlib/Algebra/MonoidAlgebra/Defs.lean b/Mathlib/Algebra/MonoidAlgebra/Defs.lean index a2574ffe76219..d7cde54dca700 100644 --- a/Mathlib/Algebra/MonoidAlgebra/Defs.lean +++ b/Mathlib/Algebra/MonoidAlgebra/Defs.lean @@ -563,7 +563,7 @@ theorem liftNC_smul [MulOneClass G] {R : Type*} [Semiring R] (f : k →+* R) (g suffices (liftNC (↑f) g).comp (smulAddHom k (MonoidAlgebra k G) c) = (AddMonoidHom.mulLeft (f c)).comp (liftNC (↑f) g) from DFunLike.congr_fun this φ - -- Porting note: `ext` couldn't a find appropriate theorem. + -- Porting note (#11041): `ext` couldn't a find appropriate theorem. refine addHom_ext' fun a => AddMonoidHom.ext fun b => ?_ -- Porting note: `reducible` cannot be `local` so the proof gets more complex. unfold MonoidAlgebra @@ -584,7 +584,7 @@ variable (k) [Semiring k] [DistribSMul R k] [Mul G] instance isScalarTower_self [IsScalarTower R k k] : IsScalarTower R (MonoidAlgebra k G) (MonoidAlgebra k G) := ⟨fun t a b => by - -- Porting note: `ext` → `refine Finsupp.ext fun _ => ?_` + -- Porting note (#11041): `ext` → `refine Finsupp.ext fun _ => ?_` refine Finsupp.ext fun m => ?_ -- Porting note: `refine` & `rw` are required because `simp` behaves differently. classical @@ -600,7 +600,7 @@ also commute with the algebra multiplication. -/ instance smulCommClass_self [SMulCommClass R k k] : SMulCommClass R (MonoidAlgebra k G) (MonoidAlgebra k G) := ⟨fun t a b => by - -- Porting note: `ext` → `refine Finsupp.ext fun _ => ?_` + -- Porting note (#11041): `ext` → `refine Finsupp.ext fun _ => ?_` refine Finsupp.ext fun m => ?_ -- Porting note: `refine` & `rw` are required because `simp` behaves differently. classical @@ -742,7 +742,7 @@ protected noncomputable def opRingEquiv [Monoid G] : rw [← AddEquiv.coe_toAddMonoidHom] refine Iff.mpr (AddMonoidHom.map_mul_iff (R := (MonoidAlgebra k G)ᵐᵒᵖ) (S := MonoidAlgebra kᵐᵒᵖ Gᵐᵒᵖ) _) ?_ - -- Porting note: Was `ext`. + -- Porting note (#11041): Was `ext`. refine AddMonoidHom.mul_op_ext _ _ <| addHom_ext' fun i₁ => AddMonoidHom.ext fun r₁ => AddMonoidHom.mul_op_ext _ _ <| addHom_ext' fun i₂ => AddMonoidHom.ext fun r₂ => ?_ -- Porting note: `reducible` cannot be `local` so proof gets long. @@ -924,7 +924,7 @@ instance nonUnitalNonAssocSemiring : NonUnitalNonAssocSemiring k[G] := simp only [mul_def] exact Eq.trans (congr_arg (sum f) (funext₂ fun a₁ b₁ => sum_zero_index)) sum_zero nsmul := fun n f => n • f - -- Porting note: `ext` → `refine Finsupp.ext fun _ => ?_` + -- Porting note (#11041): `ext` → `refine Finsupp.ext fun _ => ?_` nsmul_zero := by intros refine Finsupp.ext fun _ => ?_ @@ -1399,7 +1399,7 @@ protected noncomputable def opRingEquiv [AddCommMonoid G] : rw [Equiv.toFun_as_coe, AddEquiv.toEquiv_eq_coe]; erw [AddEquiv.coe_toEquiv] rw [← AddEquiv.coe_toAddMonoidHom] refine Iff.mpr (AddMonoidHom.map_mul_iff (R := k[G]ᵐᵒᵖ) (S := kᵐᵒᵖ[G]) _) ?_ - -- Porting note: Was `ext`. + -- Porting note (#11041): Was `ext`. refine AddMonoidHom.mul_op_ext _ _ <| addHom_ext' fun i₁ => AddMonoidHom.ext fun r₁ => AddMonoidHom.mul_op_ext _ _ <| addHom_ext' fun i₂ => AddMonoidHom.ext fun r₂ => ?_ -- Porting note: `reducible` cannot be `local` so proof gets long. diff --git a/Mathlib/Algebra/MonoidAlgebra/Division.lean b/Mathlib/Algebra/MonoidAlgebra/Division.lean index 3eb1337dc2017..42c2c79cdba87 100644 --- a/Mathlib/Algebra/MonoidAlgebra/Division.lean +++ b/Mathlib/Algebra/MonoidAlgebra/Division.lean @@ -69,14 +69,14 @@ theorem zero_divOf (g : G) : (0 : k[G]) /ᵒᶠ g = 0 := @[simp] theorem divOf_zero (x : k[G]) : x /ᵒᶠ 0 = x := by - refine Finsupp.ext fun _ => ?_ -- Porting note: `ext` doesn't work + refine Finsupp.ext fun _ => ?_ -- Porting note (#11041): `ext` doesn't work simp only [AddMonoidAlgebra.divOf_apply, zero_add] theorem add_divOf (x y : k[G]) (g : G) : (x + y) /ᵒᶠ g = x /ᵒᶠ g + y /ᵒᶠ g := map_add (Finsupp.comapDomain.addMonoidHom _) _ _ theorem divOf_add (x : k[G]) (a b : G) : x /ᵒᶠ (a + b) = x /ᵒᶠ a /ᵒᶠ b := by - refine Finsupp.ext fun _ => ?_ -- Porting note: `ext` doesn't work + refine Finsupp.ext fun _ => ?_ -- Porting note (#11041): `ext` doesn't work simp only [AddMonoidAlgebra.divOf_apply, add_assoc] /-- A bundled version of `AddMonoidAlgebra.divOf`. -/ @@ -93,13 +93,13 @@ noncomputable def divOfHom : Multiplicative G →* AddMonoid.End k[G] where (divOf_add _ _ _) theorem of'_mul_divOf (a : G) (x : k[G]) : of' k G a * x /ᵒᶠ a = x := by - refine Finsupp.ext fun _ => ?_ -- Porting note: `ext` doesn't work + refine Finsupp.ext fun _ => ?_ -- Porting note (#11041): `ext` doesn't work rw [AddMonoidAlgebra.divOf_apply, of'_apply, single_mul_apply_aux, one_mul] intro c exact add_right_inj _ theorem mul_of'_divOf (x : k[G]) (a : G) : x * of' k G a /ᵒᶠ a = x := by - refine Finsupp.ext fun _ => ?_ -- Porting note: `ext` doesn't work + refine Finsupp.ext fun _ => ?_ -- Porting note (#11041): `ext` doesn't work rw [AddMonoidAlgebra.divOf_apply, of'_apply, mul_single_apply_aux, mul_one] intro c rw [add_comm] @@ -134,14 +134,14 @@ theorem modOf_apply_self_add (x : k[G]) (g : G) (d : G) : (x %ᵒᶠ g) (g + d) modOf_apply_of_exists_add _ _ _ ⟨_, rfl⟩ theorem of'_mul_modOf (g : G) (x : k[G]) : of' k G g * x %ᵒᶠ g = 0 := by - refine Finsupp.ext fun g' => ?_ -- Porting note: `ext g'` doesn't work + refine Finsupp.ext fun g' => ?_ -- Porting note (#11041): `ext g'` doesn't work rw [Finsupp.zero_apply] obtain ⟨d, rfl⟩ | h := em (∃ d, g' = g + d) · rw [modOf_apply_self_add] · rw [modOf_apply_of_not_exists_add _ _ _ h, of'_apply, single_mul_apply_of_not_exists_add _ _ h] theorem mul_of'_modOf (x : k[G]) (g : G) : x * of' k G g %ᵒᶠ g = 0 := by - refine Finsupp.ext fun g' => ?_ -- Porting note: `ext g'` doesn't work + refine Finsupp.ext fun g' => ?_ -- Porting note (#11041): `ext g'` doesn't work rw [Finsupp.zero_apply] obtain ⟨d, rfl⟩ | h := em (∃ d, g' = g + d) · rw [modOf_apply_self_add] @@ -153,7 +153,7 @@ theorem of'_modOf (g : G) : of' k G g %ᵒᶠ g = 0 := by theorem divOf_add_modOf (x : k[G]) (g : G) : of' k G g * (x /ᵒᶠ g) + x %ᵒᶠ g = x := by - refine Finsupp.ext fun g' => ?_ -- Porting note: `ext` doesn't work + refine Finsupp.ext fun g' => ?_ -- Porting note (#11041): `ext` doesn't work rw [Finsupp.add_apply] -- Porting note: changed from `simp_rw` which can't see through the type obtain ⟨d, rfl⟩ | h := em (∃ d, g' = g + d) swap diff --git a/Mathlib/Algebra/MonoidAlgebra/ToDirectSum.lean b/Mathlib/Algebra/MonoidAlgebra/ToDirectSum.lean index aa57dcae344e6..9f6803f7cbf49 100644 --- a/Mathlib/Algebra/MonoidAlgebra/ToDirectSum.lean +++ b/Mathlib/Algebra/MonoidAlgebra/ToDirectSum.lean @@ -129,7 +129,7 @@ theorem toDirectSum_mul [DecidableEq ι] [AddMonoid ι] [Semiring M] (f g : AddM let _ : NonUnitalNonAssocSemiring (ι →₀ M) := AddMonoidAlgebra.nonUnitalNonAssocSemiring revert f g rw [AddMonoidHom.map_mul_iff] - -- Porting note: does not find `addHom_ext'`, was `ext (xi xv yi yv) : 4` + -- Porting note (#11041): does not find `addHom_ext'`, was `ext (xi xv yi yv) : 4` refine Finsupp.addHom_ext' fun xi => AddMonoidHom.ext fun xv => ?_ refine Finsupp.addHom_ext' fun yi => AddMonoidHom.ext fun yv => ?_ dsimp only [AddMonoidHom.comp_apply, AddMonoidHom.compl₂_apply, AddMonoidHom.compr₂_apply, diff --git a/Mathlib/Algebra/MvPolynomial/Basic.lean b/Mathlib/Algebra/MvPolynomial/Basic.lean index 1c4137397a564..c5e9e3036cee4 100644 --- a/Mathlib/Algebra/MvPolynomial/Basic.lean +++ b/Mathlib/Algebra/MvPolynomial/Basic.lean @@ -404,12 +404,12 @@ theorem induction_on {M : MvPolynomial σ R → Prop} (p : MvPolynomial σ R) (h theorem ringHom_ext {A : Type*} [Semiring A] {f g : MvPolynomial σ R →+* A} (hC : ∀ r, f (C r) = g (C r)) (hX : ∀ i, f (X i) = g (X i)) : f = g := by refine AddMonoidAlgebra.ringHom_ext' ?_ ?_ - -- Porting note: this has high priority, but Lean still chooses `RingHom.ext`, why? + -- Porting note (#11041): this has high priority, but Lean still chooses `RingHom.ext`, why? -- probably because of the type synonym · ext x exact hC _ · apply Finsupp.mulHom_ext'; intros x - -- Porting note: `Finsupp.mulHom_ext'` needs to have increased priority + -- Porting note (#11041): `Finsupp.mulHom_ext'` needs to have increased priority apply MonoidHom.ext_mnat exact hX _ diff --git a/Mathlib/Algebra/Polynomial/Laurent.lean b/Mathlib/Algebra/Polynomial/Laurent.lean index cd5d3cd5dc75a..c139a6417371f 100644 --- a/Mathlib/Algebra/Polynomial/Laurent.lean +++ b/Mathlib/Algebra/Polynomial/Laurent.lean @@ -87,7 +87,7 @@ scoped[LaurentPolynomial] notation:9000 R "[T;T⁻¹]" => LaurentPolynomial R open LaurentPolynomial --- Porting note: `ext` no longer applies `Finsupp.ext` automatically +-- Porting note (#5229): `ext` no longer applies `Finsupp.ext` automatically @[ext] theorem LaurentPolynomial.ext [Semiring R] {p q : R[T;T⁻¹]} (h : ∀ a, p a = q a) : p = q := Finsupp.ext h diff --git a/Mathlib/AlgebraicGeometry/Restrict.lean b/Mathlib/AlgebraicGeometry/Restrict.lean index db1fe2f6dad5c..3eadf0265e3fc 100644 --- a/Mathlib/AlgebraicGeometry/Restrict.lean +++ b/Mathlib/AlgebraicGeometry/Restrict.lean @@ -392,13 +392,13 @@ theorem image_morphismRestrict_preimage {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y. -- Porting note: this rewrite was not necessary rw [SetLike.mem_coe] convert hx' - -- Porting note: `ext1` is not compiling + -- Porting note (#11041): `ext1` is not compiling refine Subtype.ext ?_ exact (morphismRestrict_base_coe f U ⟨x, hx⟩).symm · rintro ⟨⟨x, hx⟩, hx' : _ ∈ V.1, rfl : x = _⟩ refine ⟨⟨_, hx⟩, (?_ : (f ∣_ U).base ⟨x, hx⟩ ∈ V.1), rfl⟩ convert hx' - -- Porting note: `ext1` is compiling + -- Porting note (#11041): `ext1` is compiling refine Subtype.ext ?_ exact morphismRestrict_base_coe f U ⟨x, hx⟩ diff --git a/Mathlib/AlgebraicGeometry/Spec.lean b/Mathlib/AlgebraicGeometry/Spec.lean index 349fe5db9ccec..25719eb7d693f 100644 --- a/Mathlib/AlgebraicGeometry/Spec.lean +++ b/Mathlib/AlgebraicGeometry/Spec.lean @@ -287,7 +287,7 @@ instance isIso_toSpecΓ (R : CommRingCat.{u}) : IsIso (toSpecΓ R) := by @[reassoc] theorem Spec_Γ_naturality {R S : CommRingCat.{u}} (f : R ⟶ S) : f ≫ toSpecΓ S = toSpecΓ R ≫ Γ.map (Spec.toLocallyRingedSpace.map f.op).op := by - -- Porting note: `ext` failed to pick up one of the three lemmas + -- Porting note (#11041): `ext` failed to pick up one of the three lemmas refine RingHom.ext fun x => Subtype.ext <| funext fun x' => ?_; symm apply Localization.localRingHom_to_map diff --git a/Mathlib/AlgebraicTopology/SimplexCategory.lean b/Mathlib/AlgebraicTopology/SimplexCategory.lean index 036111e7e6ab6..ac2101ce20054 100644 --- a/Mathlib/AlgebraicTopology/SimplexCategory.lean +++ b/Mathlib/AlgebraicTopology/SimplexCategory.lean @@ -142,7 +142,7 @@ lemma id_toOrderHom (a : SimplexCategory) : lemma comp_toOrderHom {a b c : SimplexCategory} (f : a ⟶ b) (g : b ⟶ c) : (f ≫ g).toOrderHom = g.toOrderHom.comp f.toOrderHom := rfl --- Porting note: added because `Hom.ext'` is not triggered automatically +-- Porting note (#5229): added because `Hom.ext'` is not triggered automatically @[ext] theorem Hom.ext {a b : SimplexCategory} (f g : a ⟶ b) : f.toOrderHom = g.toOrderHom → f = g := diff --git a/Mathlib/AlgebraicTopology/SplitSimplicialObject.lean b/Mathlib/AlgebraicTopology/SplitSimplicialObject.lean index 246867b17fe5c..48c6023dd1053 100644 --- a/Mathlib/AlgebraicTopology/SplitSimplicialObject.lean +++ b/Mathlib/AlgebraicTopology/SplitSimplicialObject.lean @@ -351,7 +351,7 @@ variable {C} namespace Split --- Porting note: added as `Hom.ext` is not triggered automatically +-- Porting note (#5229): added as `Hom.ext` is not triggered automatically @[ext] theorem hom_ext {S₁ S₂ : Split C} (Φ₁ Φ₂ : S₁ ⟶ S₂) (h : ∀ n : ℕ, Φ₁.f n = Φ₂.f n) : Φ₁ = Φ₂ := Hom.ext _ _ h diff --git a/Mathlib/CategoryTheory/Abelian/Pseudoelements.lean b/Mathlib/CategoryTheory/Abelian/Pseudoelements.lean index b472107dc47f5..7c2ed53010962 100644 --- a/Mathlib/CategoryTheory/Abelian/Pseudoelements.lean +++ b/Mathlib/CategoryTheory/Abelian/Pseudoelements.lean @@ -258,7 +258,7 @@ theorem zero_morphism_ext {P Q : C} (f : P ⟶ Q) : (∀ a, f a = 0) → f = 0 : theorem zero_morphism_ext' {P Q : C} (f : P ⟶ Q) : (∀ a, f a = 0) → 0 = f := Eq.symm ∘ zero_morphism_ext f --- Porting note: these are no longer valid as `ext` lemmas. +-- Porting note (#11182): these are no longer valid as `ext` lemmas. -- scoped[Pseudoelement] -- attribute [ext] -- CategoryTheory.Abelian.Pseudoelement.zero_morphism_ext diff --git a/Mathlib/CategoryTheory/Bicategory/NaturalTransformation/Oplax.lean b/Mathlib/CategoryTheory/Bicategory/NaturalTransformation/Oplax.lean index 28abd4f9c4360..2d6217112669a 100644 --- a/Mathlib/CategoryTheory/Bicategory/NaturalTransformation/Oplax.lean +++ b/Mathlib/CategoryTheory/Bicategory/NaturalTransformation/Oplax.lean @@ -245,7 +245,7 @@ instance category (F G : OplaxFunctor B C) : Category (F ⟶ G) where id := Modification.id comp := Modification.vcomp --- Porting note: duplicating the `ext` lemma. +-- Porting note (#5229): duplicating the `ext` lemma. @[ext] lemma ext {F G : OplaxFunctor B C} {α β : F ⟶ G} {m n : α ⟶ β} (w : ∀ b, m.app b = n.app b) : m = n := by diff --git a/Mathlib/CategoryTheory/Closed/Functor.lean b/Mathlib/CategoryTheory/Closed/Functor.lean index c0348d31a52bf..a725454594c48 100644 --- a/Mathlib/CategoryTheory/Closed/Functor.lean +++ b/Mathlib/CategoryTheory/Closed/Functor.lean @@ -79,14 +79,14 @@ theorem expComparison_ev (A B : C) : Limits.prod.map (𝟙 (F.obj A)) ((expComparison F A).app B) ≫ (exp.ev (F.obj A)).app (F.obj B) = inv (prodComparison F _ _) ≫ F.map ((exp.ev _).app _) := by convert mateEquiv_counit _ _ (prodComparisonNatIso F A).inv B using 2 - apply IsIso.inv_eq_of_hom_inv_id -- Porting note: was `ext` + apply IsIso.inv_eq_of_hom_inv_id -- Porting note (#11041): was `ext` simp only [Limits.prodComparisonNatIso_inv, asIso_inv, NatIso.isIso_inv_app, IsIso.hom_inv_id] theorem coev_expComparison (A B : C) : F.map ((exp.coev A).app B) ≫ (expComparison F A).app (A ⨯ B) = (exp.coev _).app (F.obj B) ≫ (exp (F.obj A)).map (inv (prodComparison F A B)) := by convert unit_mateEquiv _ _ (prodComparisonNatIso F A).inv B using 3 - apply IsIso.inv_eq_of_hom_inv_id -- Porting note: was `ext` + apply IsIso.inv_eq_of_hom_inv_id -- Porting note (#11041): was `ext` dsimp simp diff --git a/Mathlib/CategoryTheory/Comma/Basic.lean b/Mathlib/CategoryTheory/Comma/Basic.lean index 8cff758d67463..74eebe08f5cf9 100644 --- a/Mathlib/CategoryTheory/Comma/Basic.lean +++ b/Mathlib/CategoryTheory/Comma/Basic.lean @@ -106,7 +106,7 @@ section variable {X Y Z : Comma L R} {f : X ⟶ Y} {g : Y ⟶ Z} --- Porting note: this lemma was added because `CommaMorphism.ext` +-- Porting note (#5229): this lemma was added because `CommaMorphism.ext` -- was not triggered automatically @[ext] lemma hom_ext (f g : X ⟶ Y) (h₁ : f.left = g.left) (h₂ : f.right = g.right) : f = g := diff --git a/Mathlib/CategoryTheory/Elements.lean b/Mathlib/CategoryTheory/Elements.lean index 3ff60562fdb11..588561a7941f8 100644 --- a/Mathlib/CategoryTheory/Elements.lean +++ b/Mathlib/CategoryTheory/Elements.lean @@ -47,7 +47,7 @@ def Functor.Elements (F : C ⥤ Type w) := /-- Constructor for the type `F.Elements` when `F` is a functor to types. -/ abbrev Functor.elementsMk (F : C ⥤ Type w) (X : C) (x : F.obj X) : F.Elements := ⟨X, x⟩ --- Porting note: added because Sigma.ext would be triggered automatically +-- Porting note (#5229): added because Sigma.ext would be triggered automatically lemma Functor.Elements.ext {F : C ⥤ Type w} (x y : F.Elements) (h₁ : x.fst = y.fst) (h₂ : F.map (eqToHom h₁) x.snd = y.snd) : x = y := by cases x diff --git a/Mathlib/CategoryTheory/Endomorphism.lean b/Mathlib/CategoryTheory/Endomorphism.lean index 1e553ff57a4dc..4845f0a202a03 100644 --- a/Mathlib/CategoryTheory/Endomorphism.lean +++ b/Mathlib/CategoryTheory/Endomorphism.lean @@ -114,7 +114,7 @@ def Aut (X : C) := X ≅ X namespace Aut --- Porting note: added because `Iso.ext` is not triggered automatically +-- Porting note (#5229): added because `Iso.ext` is not triggered automatically @[ext] lemma ext {X : C} {φ₁ φ₂ : Aut X} (h : φ₁.hom = φ₂.hom) : φ₁ = φ₂ := Iso.ext h diff --git a/Mathlib/CategoryTheory/Functor/Category.lean b/Mathlib/CategoryTheory/Functor/Category.lean index e1f0445ab692b..e1bd5dc1cfab1 100644 --- a/Mathlib/CategoryTheory/Functor/Category.lean +++ b/Mathlib/CategoryTheory/Functor/Category.lean @@ -48,7 +48,7 @@ instance Functor.category : Category.{max u₁ v₂} (C ⥤ D) where namespace NatTrans --- Porting note: the behaviour of `ext` has changed here. +-- Porting note (#5229): the behaviour of `ext` has changed here. -- We need to provide a copy of the `NatTrans.ext` lemma, -- written in terms of `F ⟶ G` rather than `NatTrans F G`, -- or `ext` will not retrieve it from the cache. diff --git a/Mathlib/CategoryTheory/Idempotents/Karoubi.lean b/Mathlib/CategoryTheory/Idempotents/Karoubi.lean index a0cd003efe419..abcded134084e 100644 --- a/Mathlib/CategoryTheory/Idempotents/Karoubi.lean +++ b/Mathlib/CategoryTheory/Idempotents/Karoubi.lean @@ -101,7 +101,7 @@ theorem hom_ext_iff {P Q : Karoubi C} {f g : P ⟶ Q} : f = g ↔ f.f = g.f := b rw [h] · apply Hom.ext --- Porting note: added because `Hom.ext` is not triggered automatically +-- Porting note (#5229): added because `Hom.ext` is not triggered automatically @[ext] theorem hom_ext {P Q : Karoubi C} (f g : P ⟶ Q) (h : f.f = g.f) : f = g := by simpa [hom_ext_iff] using h diff --git a/Mathlib/CategoryTheory/Iso.lean b/Mathlib/CategoryTheory/Iso.lean index af07a00e52f8d..7ce50b850ef49 100644 --- a/Mathlib/CategoryTheory/Iso.lean +++ b/Mathlib/CategoryTheory/Iso.lean @@ -305,7 +305,7 @@ instance (priority := 100) mono_of_iso (f : X ⟶ Y) [IsIso f] : Mono f where rw [← Category.comp_id g, ← Category.comp_id h, ← IsIso.hom_inv_id f, ← Category.assoc, w, ← Category.assoc] --- Porting note: `@[ext]` used to accept lemmas like this. Now we add an aesop rule +-- Porting note (#11182): `@[ext]` used to accept lemmas like this. Now we add an aesop rule @[aesop apply safe (rule_sets := [CategoryTheory])] theorem inv_eq_of_hom_inv_id {f : X ⟶ Y} [IsIso f] {g : Y ⟶ X} (hom_inv_id : f ≫ g = 𝟙 X) : inv f = g := by @@ -317,7 +317,7 @@ theorem inv_eq_of_inv_hom_id {f : X ⟶ Y} [IsIso f] {g : Y ⟶ X} (inv_hom_id : apply (cancel_mono f).mp simp [inv_hom_id] --- Porting note: `@[ext]` used to accept lemmas like this. +-- Porting note (#11182): `@[ext]` used to accept lemmas like this. @[aesop apply safe (rule_sets := [CategoryTheory])] theorem eq_inv_of_hom_inv_id {f : X ⟶ Y} [IsIso f] {g : Y ⟶ X} (hom_inv_id : f ≫ g = 𝟙 X) : g = inv f := @@ -446,12 +446,12 @@ theorem isIso_of_comp_hom_eq_id (g : X ⟶ Y) [IsIso g] {f : Y ⟶ X} (h : f ≫ namespace Iso --- Porting note: `@[ext]` used to accept lemmas like this. +-- Porting note (#11182): `@[ext]` used to accept lemmas like this. @[aesop apply safe (rule_sets := [CategoryTheory])] theorem inv_ext {f : X ≅ Y} {g : Y ⟶ X} (hom_inv_id : f.hom ≫ g = 𝟙 X) : f.inv = g := ((hom_comp_eq_id f).1 hom_inv_id).symm --- Porting note: `@[ext]` used to accept lemmas like this. +-- Porting note (#11182): `@[ext]` used to accept lemmas like this. @[aesop apply safe (rule_sets := [CategoryTheory])] theorem inv_ext' {f : X ≅ Y} {g : Y ⟶ X} (hom_inv_id : f.hom ≫ g = 𝟙 X) : g = f.inv := (hom_comp_eq_id f).1 hom_inv_id diff --git a/Mathlib/CategoryTheory/Limits/Cones.lean b/Mathlib/CategoryTheory/Limits/Cones.lean index 29912d06cb23c..098f9f7adcb6a 100644 --- a/Mathlib/CategoryTheory/Limits/Cones.lean +++ b/Mathlib/CategoryTheory/Limits/Cones.lean @@ -283,7 +283,7 @@ namespace Cones /-- To give an isomorphism between cones, it suffices to give an isomorphism between their vertices which commutes with the cone maps. -/ --- Porting note: `@[ext]` used to accept lemmas like this. Now we add an aesop rule +-- Porting note (#11182): `@[ext]` used to accept lemmas like this. Now we add an aesop rule @[aesop apply safe (rule_sets := [CategoryTheory]), simps] def ext {c c' : Cone F} (φ : c.pt ≅ c'.pt) (w : ∀ j, c.π.app j = φ.hom ≫ c'.π.app j := by aesop_cat) : c ≅ c' where @@ -484,7 +484,7 @@ namespace Cocones /-- To give an isomorphism between cocones, it suffices to give an isomorphism between their vertices which commutes with the cocone maps. -/ --- Porting note: `@[ext]` used to accept lemmas like this. Now we add an aesop rule +-- Porting note (#11182): `@[ext]` used to accept lemmas like this. Now we add an aesop rule @[aesop apply safe (rule_sets := [CategoryTheory]), simps] def ext {c c' : Cocone F} (φ : c.pt ≅ c'.pt) (w : ∀ j, c.ι.app j ≫ φ.hom = c'.ι.app j := by aesop_cat) : c ≅ c' where diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean b/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean index 78e960db24953..2d16499600981 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean @@ -122,7 +122,7 @@ namespace Bicones /-- To give an isomorphism between cocones, it suffices to give an isomorphism between their vertices which commutes with the cocone maps. -/ --- Porting note: `@[ext]` used to accept lemmas like this. Now we add an aesop rule +-- Porting note (#11182): `@[ext]` used to accept lemmas like this. Now we add an aesop rule @[aesop apply safe (rule_sets := [CategoryTheory]), simps] def ext {c c' : Bicone F} (φ : c.pt ≅ c'.pt) (wι : ∀ j, c.ι j ≫ φ.hom = c'.ι j := by aesop_cat) @@ -1119,7 +1119,7 @@ namespace BinaryBicones /-- To give an isomorphism between cocones, it suffices to give an isomorphism between their vertices which commutes with the cocone maps. -/ --- Porting note: `@[ext]` used to accept lemmas like this. Now we add an aesop rule +-- Porting note (#11182): `@[ext]` used to accept lemmas like this. Now we add an aesop rule @[aesop apply safe (rule_sets := [CategoryTheory]), simps] def ext {P Q : C} {c c' : BinaryBicone P Q} (φ : c.pt ≅ c'.pt) (winl : c.inl ≫ φ.hom = c'.inl := by aesop_cat) diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Multiequalizer.lean b/Mathlib/CategoryTheory/Limits/Shapes/Multiequalizer.lean index 54c25d54c53c4..d880d204adea3 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Multiequalizer.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Multiequalizer.lean @@ -662,7 +662,7 @@ noncomputable def multicoforkEquivSigmaCofork : counitIso := NatIso.ofComponents fun K => Cofork.ext (Iso.refl _) (by - -- Porting note: in mathlib3 this was just `ext` and I don't know why it's not here + -- Porting note (#11041): in mathlib3 this was just `ext` and I don't know why it's not here apply Limits.colimit.hom_ext rintro ⟨j⟩ dsimp diff --git a/Mathlib/CategoryTheory/Monoidal/Bimod.lean b/Mathlib/CategoryTheory/Monoidal/Bimod.lean index b4118c7d293bf..b6d566f92f32a 100644 --- a/Mathlib/CategoryTheory/Monoidal/Bimod.lean +++ b/Mathlib/CategoryTheory/Monoidal/Bimod.lean @@ -120,7 +120,7 @@ instance : Category (Bimod A B) where id := id' comp f g := comp f g --- Porting note: added because `Hom.ext` is not triggered automatically +-- Porting note (#5229): added because `Hom.ext` is not triggered automatically @[ext] lemma hom_ext {M N : Bimod A B} (f g : M ⟶ N) (h : f.hom = g.hom) : f = g := Hom.ext h diff --git a/Mathlib/CategoryTheory/Monoidal/Braided/Basic.lean b/Mathlib/CategoryTheory/Monoidal/Braided/Basic.lean index ecac5bcd3be12..c0dfcace9f9a6 100644 --- a/Mathlib/CategoryTheory/Monoidal/Braided/Basic.lean +++ b/Mathlib/CategoryTheory/Monoidal/Braided/Basic.lean @@ -399,7 +399,7 @@ def comp (F : LaxBraidedFunctor C D) (G : LaxBraidedFunctor D E) : LaxBraidedFun instance categoryLaxBraidedFunctor : Category (LaxBraidedFunctor C D) := InducedCategory.category LaxBraidedFunctor.toLaxMonoidalFunctor --- Porting note: added, as `MonoidalNatTrans.ext` does not apply to morphisms. +-- Porting note (#5229): added, as `MonoidalNatTrans.ext` does not apply to morphisms. @[ext] lemma ext' {F G : LaxBraidedFunctor C D} {α β : F ⟶ G} (w : ∀ X : C, α.app X = β.app X) : α = β := MonoidalNatTrans.ext (funext w) @@ -463,7 +463,7 @@ def comp (F : BraidedFunctor C D) (G : BraidedFunctor D E) : BraidedFunctor C E instance categoryBraidedFunctor : Category (BraidedFunctor C D) := InducedCategory.category BraidedFunctor.toMonoidalFunctor --- Porting note: added, as `MonoidalNatTrans.ext` does not apply to morphisms. +-- Porting note (#5229): added, as `MonoidalNatTrans.ext` does not apply to morphisms. @[ext] lemma ext' {F G : BraidedFunctor C D} {α β : F ⟶ G} (w : ∀ X : C, α.app X = β.app X) : α = β := MonoidalNatTrans.ext (funext w) diff --git a/Mathlib/CategoryTheory/Monoidal/Center.lean b/Mathlib/CategoryTheory/Monoidal/Center.lean index 5eab3c67802ba..3e6f4d3046f5a 100644 --- a/Mathlib/CategoryTheory/Monoidal/Center.lean +++ b/Mathlib/CategoryTheory/Monoidal/Center.lean @@ -278,7 +278,7 @@ theorem associator_hom_f (X Y Z : Center C) : Hom.f (α_ X Y Z).hom = (α_ X.1 Y @[simp] theorem associator_inv_f (X Y Z : Center C) : Hom.f (α_ X Y Z).inv = (α_ X.1 Y.1 Z.1).inv := by - apply Iso.inv_ext' -- Porting note: Originally `ext` + apply Iso.inv_ext' -- Porting note (#11041): Originally `ext` rw [← associator_hom_f, ← comp_f, Iso.hom_inv_id]; rfl @[simp] @@ -287,7 +287,7 @@ theorem leftUnitor_hom_f (X : Center C) : Hom.f (λ_ X).hom = (λ_ X.1).hom := @[simp] theorem leftUnitor_inv_f (X : Center C) : Hom.f (λ_ X).inv = (λ_ X.1).inv := by - apply Iso.inv_ext' -- Porting note: Originally `ext` + apply Iso.inv_ext' -- Porting note (#11041): Originally `ext` rw [← leftUnitor_hom_f, ← comp_f, Iso.hom_inv_id]; rfl @[simp] @@ -296,7 +296,7 @@ theorem rightUnitor_hom_f (X : Center C) : Hom.f (ρ_ X).hom = (ρ_ X.1).hom := @[simp] theorem rightUnitor_inv_f (X : Center C) : Hom.f (ρ_ X).inv = (ρ_ X.1).inv := by - apply Iso.inv_ext' -- Porting note: Originally `ext` + apply Iso.inv_ext' -- Porting note (#11041): Originally `ext` rw [← rightUnitor_hom_f, ← comp_f, Iso.hom_inv_id]; rfl end diff --git a/Mathlib/CategoryTheory/Monoidal/Internal/Module.lean b/Mathlib/CategoryTheory/Monoidal/Internal/Module.lean index 4bdf23e5898f4..b579897c35ad0 100644 --- a/Mathlib/CategoryTheory/Monoidal/Internal/Module.lean +++ b/Mathlib/CategoryTheory/Monoidal/Internal/Module.lean @@ -103,7 +103,7 @@ def inverseObj (A : AlgebraCat.{u} R) : Mon_ (ModuleCat.{u} R) where one := Algebra.linearMap R A mul := LinearMap.mul' R A one_mul := by - -- Porting note: `ext` did not pick up `TensorProduct.ext` + -- Porting note (#11041): `ext` did not pick up `TensorProduct.ext` refine TensorProduct.ext <| LinearMap.ext_ring <| LinearMap.ext fun x => ?_ -- This used to be `rw`, but we need `erw` after leanprover/lean4#2644 erw [compr₂_apply, compr₂_apply, CategoryTheory.comp_apply] @@ -116,7 +116,7 @@ def inverseObj (A : AlgebraCat.{u} R) : Mon_ (ModuleCat.{u} R) where erw [LinearMap.mul'_apply, MonoidalCategory.leftUnitor_hom_apply, ← Algebra.smul_def] erw [id_apply] mul_one := by - -- Porting note: `ext` did not pick up `TensorProduct.ext` + -- Porting note (#11041): `ext` did not pick up `TensorProduct.ext` refine TensorProduct.ext <| LinearMap.ext fun x => LinearMap.ext_ring ?_ -- Porting note: this `dsimp` does nothing -- dsimp only [AlgebraCat.id_apply, TensorProduct.mk_apply, Algebra.linearMap_apply, @@ -131,7 +131,7 @@ def inverseObj (A : AlgebraCat.{u} R) : Mon_ (ModuleCat.{u} R) where erw [id_apply] mul_assoc := by set_option tactic.skipAssignedInstances false in - -- Porting note: `ext` did not pick up `TensorProduct.ext` + -- Porting note (#11041): `ext` did not pick up `TensorProduct.ext` refine TensorProduct.ext <| TensorProduct.ext <| LinearMap.ext fun x => LinearMap.ext fun y => LinearMap.ext fun z => ?_ dsimp only [AlgebraCat.id_apply, TensorProduct.mk_apply, LinearMap.compr₂_apply, @@ -176,7 +176,7 @@ def monModuleEquivalenceAlgebra : Mon_ (ModuleCat.{u} R) ≌ AlgebraCat R where map_add' := fun _ _ => rfl map_smul' := fun _ _ => rfl } mul_hom := by - -- Porting note: `ext` did not pick up `TensorProduct.ext` + -- Porting note (#11041): `ext` did not pick up `TensorProduct.ext` refine TensorProduct.ext ?_ dsimp at * rfl } @@ -186,7 +186,7 @@ def monModuleEquivalenceAlgebra : Mon_ (ModuleCat.{u} R) ≌ AlgebraCat R where map_add' := fun _ _ => rfl map_smul' := fun _ _ => rfl } mul_hom := by - -- Porting note: `ext` did not pick up `TensorProduct.ext` + -- Porting note (#11041): `ext` did not pick up `TensorProduct.ext` refine TensorProduct.ext ?_ dsimp at * rfl } }) diff --git a/Mathlib/CategoryTheory/Monoidal/Mon_.lean b/Mathlib/CategoryTheory/Monoidal/Mon_.lean index be13f54b4dd13..82f033d04a045 100644 --- a/Mathlib/CategoryTheory/Monoidal/Mon_.lean +++ b/Mathlib/CategoryTheory/Monoidal/Mon_.lean @@ -104,7 +104,7 @@ instance : Category (Mon_ C) where id := id comp f g := comp f g --- Porting note: added, as `Hom.ext` does not apply to a morphism. +-- Porting note (#5229): added, as `Hom.ext` does not apply to a morphism. @[ext] lemma ext {X Y : Mon_ C} {f g : X ⟶ Y} (w : f.hom = g.hom) : f = g := Hom.ext w diff --git a/Mathlib/CategoryTheory/Monoidal/NaturalTransformation.lean b/Mathlib/CategoryTheory/Monoidal/NaturalTransformation.lean index 37d92020f5dfd..4d73ab1b96ac5 100644 --- a/Mathlib/CategoryTheory/Monoidal/NaturalTransformation.lean +++ b/Mathlib/CategoryTheory/Monoidal/NaturalTransformation.lean @@ -81,7 +81,7 @@ theorem comp_toNatTrans_lax {F G H : LaxMonoidalFunctor C D} {α : F ⟶ G} {β instance categoryMonoidalFunctor : Category (MonoidalFunctor C D) := InducedCategory.category MonoidalFunctor.toLaxMonoidalFunctor --- Porting note: added, as `MonoidalNatTrans.ext` does not apply to morphisms. +-- Porting note (#5229): added, as `MonoidalNatTrans.ext` does not apply to morphisms. @[ext] lemma ext' {F G : LaxMonoidalFunctor C D} {α β : F ⟶ G} (w : ∀ X : C, α.app X = β.app X) : α = β := MonoidalNatTrans.ext (funext w) diff --git a/Mathlib/CategoryTheory/Pi/Basic.lean b/Mathlib/CategoryTheory/Pi/Basic.lean index 6827f96279d67..beeb60fcf3da4 100644 --- a/Mathlib/CategoryTheory/Pi/Basic.lean +++ b/Mathlib/CategoryTheory/Pi/Basic.lean @@ -49,7 +49,7 @@ theorem comp_apply {X Y Z : ∀ i, C i} (f : X ⟶ Y) (g : Y ⟶ Z) (i) : (f ≫ g : ∀ i, X i ⟶ Z i) i = f i ≫ g i := rfl --- Porting note: need to add an additional `ext` lemma. +-- Porting note (#5229): need to add an additional `ext` lemma. @[ext] lemma ext {X Y : ∀ i, C i} {f g : X ⟶ Y} (w : ∀ i, f i = g i) : f = g := funext (w ·) diff --git a/Mathlib/CategoryTheory/Sites/Pretopology.lean b/Mathlib/CategoryTheory/Sites/Pretopology.lean index 85b5f87c3c0a8..972ef7c94b08d 100644 --- a/Mathlib/CategoryTheory/Sites/Pretopology.lean +++ b/Mathlib/CategoryTheory/Sites/Pretopology.lean @@ -183,7 +183,7 @@ def trivial : Pretopology C where rcases hS g (singleton_self g) with ⟨Y, f, i, hTi⟩ refine ⟨_, f ≫ g, ?_, ?_⟩ · infer_instance - -- Porting note: the next four lines were just "ext (W k)" + -- Porting note (#11041): the next four lines were just "ext (W k)" apply funext rintro W apply Set.ext diff --git a/Mathlib/CategoryTheory/Sites/Sheaf.lean b/Mathlib/CategoryTheory/Sites/Sheaf.lean index 07e5279a34396..99d958c7ef686 100644 --- a/Mathlib/CategoryTheory/Sites/Sheaf.lean +++ b/Mathlib/CategoryTheory/Sites/Sheaf.lean @@ -332,7 +332,7 @@ instance instCategorySheaf : Category (Sheaf J A) where instance (X : Sheaf J A) : Inhabited (Hom X X) := ⟨𝟙 X⟩ --- Porting note: added because `Sheaf.Hom.ext` was not triggered automatically +-- Porting note (#5229): added because `Sheaf.Hom.ext` was not triggered automatically @[ext] lemma hom_ext {X Y : Sheaf J A} (x y : X ⟶ Y) (h : x.val = y.val) : x = y := Sheaf.Hom.ext h diff --git a/Mathlib/Data/PFunctor/Univariate/M.lean b/Mathlib/Data/PFunctor/Univariate/M.lean index 73fe251a4dda5..58879cd06c136 100644 --- a/Mathlib/Data/PFunctor/Univariate/M.lean +++ b/Mathlib/Data/PFunctor/Univariate/M.lean @@ -92,7 +92,7 @@ theorem truncate_eq_of_agree {n : ℕ} (x : CofixA F n) (y : CofixA F (succ n)) · -- cases' h with _ _ _ _ _ h₀ h₁ cases h simp only [truncate, Function.comp_def, eq_self_iff_true, heq_iff_eq] - -- Porting note: used to be `ext y` + -- Porting note (#11041): used to be `ext y` rename_i n_ih a f y h₁ suffices (fun x => truncate (y x)) = f by simp [this] diff --git a/Mathlib/Geometry/RingedSpace/OpenImmersion.lean b/Mathlib/Geometry/RingedSpace/OpenImmersion.lean index 0f4f4ce207163..a29fd4205a61a 100644 --- a/Mathlib/Geometry/RingedSpace/OpenImmersion.lean +++ b/Mathlib/Geometry/RingedSpace/OpenImmersion.lean @@ -123,7 +123,7 @@ noncomputable def isoRestrict : X ≅ Y.restrict H.base_open := @[reassoc (attr := simp)] theorem isoRestrict_hom_ofRestrict : (isoRestrict f).hom ≫ Y.ofRestrict _ = f := by - -- Porting note: `ext` did not pick up `NatTrans.ext` + -- Porting note (#11041): `ext` did not pick up `NatTrans.ext` refine PresheafedSpace.Hom.ext _ _ rfl <| NatTrans.ext <| funext fun x => ?_ simp only [isoRestrict_hom_c_app, NatTrans.comp_app, eqToHom_refl, ofRestrict_c_app, Category.assoc, whiskerRight_id'] @@ -330,7 +330,7 @@ def pullbackConeOfLeftFst : rfl } theorem pullback_cone_of_left_condition : pullbackConeOfLeftFst f g ≫ f = Y.ofRestrict _ ≫ g := by - -- Porting note: `ext` did not pick up `NatTrans.ext` + -- Porting note (#11041): `ext` did not pick up `NatTrans.ext` refine PresheafedSpace.Hom.ext _ _ ?_ <| NatTrans.ext <| funext fun U => ?_ · simpa using pullback.condition · induction U using Opposite.rec' @@ -388,7 +388,7 @@ def pullbackConeOfLeftLift : s.pt ⟶ (pullbackConeOfLeft f g).pt where -- this lemma is not a `simp` lemma, because it is an implementation detail theorem pullbackConeOfLeftLift_fst : pullbackConeOfLeftLift f g s ≫ (pullbackConeOfLeft f g).fst = s.fst := by - -- Porting note: `ext` did not pick up `NatTrans.ext` + -- Porting note (#11041): `ext` did not pick up `NatTrans.ext` refine PresheafedSpace.Hom.ext _ _ ?_ <| NatTrans.ext <| funext fun x => ?_ · change pullback.lift _ _ _ ≫ pullback.fst _ _ = _ simp @@ -407,7 +407,7 @@ theorem pullbackConeOfLeftLift_fst : -- this lemma is not a `simp` lemma, because it is an implementation detail theorem pullbackConeOfLeftLift_snd : pullbackConeOfLeftLift f g s ≫ (pullbackConeOfLeft f g).snd = s.snd := by - -- Porting note: `ext` did not pick up `NatTrans.ext` + -- Porting note (#11041): `ext` did not pick up `NatTrans.ext` refine PresheafedSpace.Hom.ext _ _ ?_ <| NatTrans.ext <| funext fun x => ?_ · change pullback.lift _ _ _ ≫ pullback.snd _ _ = _ simp diff --git a/Mathlib/Geometry/RingedSpace/PresheafedSpace/HasColimits.lean b/Mathlib/Geometry/RingedSpace/PresheafedSpace/HasColimits.lean index ce942b2c84104..9b2339708fe4b 100644 --- a/Mathlib/Geometry/RingedSpace/PresheafedSpace/HasColimits.lean +++ b/Mathlib/Geometry/RingedSpace/PresheafedSpace/HasColimits.lean @@ -224,7 +224,7 @@ theorem desc_fac (F : J ⥤ PresheafedSpace.{_, _, v} C) (s : Cocone F) (j : J) (colimitCocone F).ι.app j ≫ desc F s = s.ι.app j := by ext U · simp [desc] - · -- Porting note: the original proof is just `ext; dsimp [desc, descCApp]; simpa`, + · -- Porting note (#11041): the original proof is just `ext; dsimp [desc, descCApp]; simpa`, -- but this has to be expanded a bit rw [NatTrans.comp_app, PresheafedSpace.comp_c_app, whiskerRight_app] dsimp [desc, descCApp] diff --git a/Mathlib/Order/Category/OmegaCompletePartialOrder.lean b/Mathlib/Order/Category/OmegaCompletePartialOrder.lean index 7a02ed603f9c1..7f8ab62dfc249 100644 --- a/Mathlib/Order/Category/OmegaCompletePartialOrder.lean +++ b/Mathlib/Order/Category/OmegaCompletePartialOrder.lean @@ -81,7 +81,7 @@ def isProduct (J : Type v) (f : J → ωCPO) : IsLimit (product f) where ⟨⟨fun t j => (s.π.app ⟨j⟩).toFun t, fun _ _ h j => (s.π.app ⟨j⟩).monotone h⟩, fun x => funext fun j => (s.π.app ⟨j⟩).continuous x⟩ uniq s m w := by - ext t; funext j -- Porting note: Originally `ext t j` + ext t; funext j -- Porting note (#11041): Originally `ext t j` change m.toFun t j = (s.π.app ⟨j⟩).toFun t rw [← w ⟨j⟩] rfl @@ -98,7 +98,7 @@ instance omegaCompletePartialOrderEqualizer {α β : Type*} [OmegaCompletePartia OmegaCompletePartialOrder.subtype _ fun c hc => by rw [f.continuous, g.continuous] congr 1 - apply OrderHom.ext; funext x -- Porting note: Originally `ext` + apply OrderHom.ext; funext x -- Porting note (#11041): Originally `ext` apply hc _ ⟨_, rfl⟩ namespace HasEqualizers @@ -122,7 +122,7 @@ def isEqualizer {X Y : ωCPO.{v}} (f g : X ⟶ Y) : IsLimit (equalizer f g) := monotone' := fun _ _ h => s.ι.monotone h map_ωSup' := fun x => Subtype.ext (s.ι.continuous x) }, by ext; rfl, fun hm => by - apply ContinuousHom.ext _ _ fun x => Subtype.ext ?_ -- Porting note: Originally `ext` + apply ContinuousHom.ext _ _ fun x => Subtype.ext ?_ -- Porting note (#11041): Originally `ext` apply ContinuousHom.congr_fun hm⟩ end HasEqualizers diff --git a/Mathlib/RepresentationTheory/Action/Basic.lean b/Mathlib/RepresentationTheory/Action/Basic.lean index 9097a7d47988b..4c1ae310a5812 100644 --- a/Mathlib/RepresentationTheory/Action/Basic.lean +++ b/Mathlib/RepresentationTheory/Action/Basic.lean @@ -111,7 +111,7 @@ instance : Category (Action V G) where id M := Hom.id M comp f g := Hom.comp f g --- Porting note: added because `Hom.ext` is not triggered automatically +-- Porting note (#5229): added because `Hom.ext` is not triggered automatically @[ext] lemma hom_ext {M N : Action V G} (φ₁ φ₂ : M ⟶ N) (h : φ₁.hom = φ₂.hom) : φ₁ = φ₂ := Hom.ext h diff --git a/Mathlib/RingTheory/DedekindDomain/FiniteAdeleRing.lean b/Mathlib/RingTheory/DedekindDomain/FiniteAdeleRing.lean index 5ad19cc5ca6ae..6481e7153138a 100644 --- a/Mathlib/RingTheory/DedekindDomain/FiniteAdeleRing.lean +++ b/Mathlib/RingTheory/DedekindDomain/FiniteAdeleRing.lean @@ -106,7 +106,7 @@ def Coe.addMonoidHom : AddMonoidHom (R_hat R K) (K_hat R K) where toFun := (↑) map_zero' := rfl map_add' x y := by - -- Porting note: was `ext v` + -- Porting note (#11041): was `ext v` refine funext fun v => ?_ simp only [coe_apply, Pi.add_apply, Subring.coe_add] -- Porting note: added @@ -119,7 +119,7 @@ def Coe.ringHom : RingHom (R_hat R K) (K_hat R K) := toFun := (↑) map_one' := rfl map_mul' := fun x y => by - -- Porting note: was `ext p` + -- Porting note (#11041): was `ext p` refine funext fun p => ?_ simp only [Pi.mul_apply, Subring.coe_mul] -- Porting note: added diff --git a/Mathlib/Topology/Category/CompHaus/Basic.lean b/Mathlib/Topology/Category/CompHaus/Basic.lean index 0306bef4b8f7d..00e0ff959c9f6 100644 --- a/Mathlib/Topology/Category/CompHaus/Basic.lean +++ b/Mathlib/Topology/Category/CompHaus/Basic.lean @@ -93,7 +93,7 @@ noncomputable def stoneCechEquivalence (X : TopCat.{u}) (Y : CompHaus.{u}) : continuous_toFun := continuous_stoneCechExtend f.2 } left_inv := by rintro ⟨f : StoneCech X ⟶ Y, hf : Continuous f⟩ - -- Porting note: `ext` fails. + -- Porting note (#11041): `ext` fails. apply ContinuousMap.ext intro (x : StoneCech X) refine congr_fun ?_ x @@ -103,7 +103,7 @@ noncomputable def stoneCechEquivalence (X : TopCat.{u}) (Y : CompHaus.{u}) : apply continuous_stoneCechUnit right_inv := by rintro ⟨f : (X : Type _) ⟶ Y, hf : Continuous f⟩ - -- Porting note: `ext` fails. + -- Porting note (#11041): `ext` fails. apply ContinuousMap.ext intro exact congr_fun (stoneCechExtend_extends hf) _ @@ -203,7 +203,7 @@ theorem epi_iff_surjective {X Y : CompHaus.{u}} (f : X ⟶ Y) : Epi f ↔ Functi have H : h = g := by rw [← cancel_epi f] ext x - -- Porting note: `ext` doesn't apply these two lemmas. + -- Porting note (#11041): `ext` doesn't apply these two lemmas. apply ULift.ext apply Subtype.ext dsimp diff --git a/Mathlib/Topology/Category/Profinite/AsLimit.lean b/Mathlib/Topology/Category/Profinite/AsLimit.lean index 812192fbc0757..9a23eb81a082a 100644 --- a/Mathlib/Topology/Category/Profinite/AsLimit.lean +++ b/Mathlib/Topology/Category/Profinite/AsLimit.lean @@ -66,7 +66,8 @@ instance isIso_asLimitCone_lift : IsIso ((limitConeIsLimit.{u, u} X.diagram).lif · obtain ⟨b, hb⟩ := DiscreteQuotient.exists_of_compat (fun S => a.val S) fun _ _ h => a.prop (homOfLE h) use b - -- ext S : 3 -- Porting note: `ext` does not work, replaced with following three lines. + -- ext S : 3 -- Porting note (#11041): `ext` does not work, replaced with following + -- three lines. apply Subtype.ext apply funext rintro S diff --git a/Mathlib/Topology/Category/TopCat/Limits/Products.lean b/Mathlib/Topology/Category/TopCat/Limits/Products.lean index 4a6576bc290e8..cd45197d42a40 100644 --- a/Mathlib/Topology/Category/TopCat/Limits/Products.lean +++ b/Mathlib/Topology/Category/TopCat/Limits/Products.lean @@ -161,7 +161,7 @@ def prodBinaryFanIsLimit (X Y : TopCat.{u}) : IsLimit (prodBinaryFan X Y) where rintro S (_ | _) <;> {dsimp; ext; rfl} uniq := by intro S m h - -- Porting note: used to be `ext x` + -- Porting note (#11041): used to be `ext x` refine ContinuousMap.ext (fun (x : ↥(S.pt)) => Prod.ext ?_ ?_) · specialize h ⟨WalkingPair.left⟩ apply_fun fun e => e x at h diff --git a/Mathlib/Topology/Sheaves/PresheafOfFunctions.lean b/Mathlib/Topology/Sheaves/PresheafOfFunctions.lean index ce0908cb25970..4ce4eaa2a3fd1 100644 --- a/Mathlib/Topology/Sheaves/PresheafOfFunctions.lean +++ b/Mathlib/Topology/Sheaves/PresheafOfFunctions.lean @@ -121,7 +121,7 @@ this is a ring homomorphism (with respect to the pointwise ring operations on fu def map (X : TopCat.{u}ᵒᵖ) {R S : TopCommRingCat.{u}} (φ : R ⟶ S) : continuousFunctions X R ⟶ continuousFunctions X S where toFun g := g ≫ (forget₂ TopCommRingCat TopCat).map φ - -- Porting note: `ext` tactic does not work, since Lean can't see through `R ⟶ S` is just + -- Porting note (#11041): `ext` tactic does not work, since Lean can't see through `R ⟶ S` is just -- continuous ring homomorphism map_one' := ContinuousMap.ext fun _ => φ.1.map_one map_zero' := ContinuousMap.ext fun _ => φ.1.map_zero From bfe51d4ef05f50660a2ffb13e0a4286ac7f3a5bb Mon Sep 17 00:00:00 2001 From: damiano Date: Wed, 16 Oct 2024 15:23:31 +0000 Subject: [PATCH 033/505] chore: yet more unused variables (#17826) --- Mathlib/AlgebraicTopology/FundamentalGroupoid/Basic.lean | 4 ++-- .../FundamentalGroupoid/FundamentalGroup.lean | 4 ++-- Mathlib/CategoryTheory/Limits/Shapes/Pullback/Assoc.lean | 2 +- Mathlib/CategoryTheory/Sites/Canonical.lean | 4 +--- Mathlib/Computability/Reduce.lean | 7 ++----- Mathlib/Control/Fold.lean | 1 - Mathlib/Control/Functor/Multivariate.lean | 2 +- Mathlib/Data/BitVec.lean | 2 +- Mathlib/Data/DFinsupp/Multiset.lean | 4 ++-- Mathlib/Data/TypeVec.lean | 4 ++-- Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean | 4 ++-- Mathlib/MeasureTheory/Measure/Hausdorff.lean | 4 ++-- Mathlib/MeasureTheory/Measure/LevyProkhorovMetric.lean | 6 +++--- Mathlib/MeasureTheory/Order/UpperLower.lean | 2 +- Mathlib/ModelTheory/Complexity.lean | 7 ++----- Mathlib/ModelTheory/ElementarySubstructures.lean | 3 +-- Mathlib/ModelTheory/Encoding.lean | 5 ++--- Mathlib/ModelTheory/Syntax.lean | 3 +-- Mathlib/RingTheory/GradedAlgebra/HomogeneousIdeal.lean | 2 +- Mathlib/RingTheory/Localization/Finiteness.lean | 2 +- Mathlib/RingTheory/TensorProduct/MvPolynomial.lean | 5 ++--- Mathlib/Topology/MetricSpace/Holder.lean | 2 +- Mathlib/Topology/MetricSpace/MetricSeparated.lean | 2 +- Mathlib/Topology/Order/DenselyOrdered.lean | 2 +- Mathlib/Topology/Order/IntermediateValue.lean | 6 ++---- Mathlib/Topology/Order/IsLUB.lean | 5 ++--- Mathlib/Topology/Order/Monotone.lean | 7 +++---- Mathlib/Topology/UniformSpace/Equicontinuity.lean | 4 ++-- .../Topology/UniformSpace/UniformConvergenceTopology.lean | 4 ++-- 29 files changed, 46 insertions(+), 63 deletions(-) diff --git a/Mathlib/AlgebraicTopology/FundamentalGroupoid/Basic.lean b/Mathlib/AlgebraicTopology/FundamentalGroupoid/Basic.lean index 2af55680d6b61..441b8fdab6dd2 100644 --- a/Mathlib/AlgebraicTopology/FundamentalGroupoid/Basic.lean +++ b/Mathlib/AlgebraicTopology/FundamentalGroupoid/Basic.lean @@ -20,9 +20,9 @@ group of `x`. open CategoryTheory -universe u v +universe u -variable {X : Type u} {Y : Type v} [TopologicalSpace X] [TopologicalSpace Y] +variable {X : Type u} [TopologicalSpace X] variable {x₀ x₁ : X} noncomputable section diff --git a/Mathlib/AlgebraicTopology/FundamentalGroupoid/FundamentalGroup.lean b/Mathlib/AlgebraicTopology/FundamentalGroupoid/FundamentalGroup.lean index b1c4b82e4469b..e5d2e226cd519 100644 --- a/Mathlib/AlgebraicTopology/FundamentalGroupoid/FundamentalGroup.lean +++ b/Mathlib/AlgebraicTopology/FundamentalGroupoid/FundamentalGroup.lean @@ -17,9 +17,9 @@ of `x` i.e. the group with elements being loops based at `x` (quotiented by homo -/ -universe u v +universe u -variable {X : Type u} {Y : Type v} [TopologicalSpace X] [TopologicalSpace Y] +variable {X : Type u} [TopologicalSpace X] variable {x₀ x₁ : X} noncomputable section diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Pullback/Assoc.lean b/Mathlib/CategoryTheory/Limits/Shapes/Pullback/Assoc.lean index 541f09bc26155..795528a7feb85 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Pullback/Assoc.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Pullback/Assoc.lean @@ -21,7 +21,7 @@ universe w v₁ v₂ v u u₂ namespace CategoryTheory.Limits -variable {C : Type u} [Category.{v} C] {W X Y Z : C} +variable {C : Type u} [Category.{v} C] section PullbackAssoc diff --git a/Mathlib/CategoryTheory/Sites/Canonical.lean b/Mathlib/CategoryTheory/Sites/Canonical.lean index ec5fdd447f5e9..a31fc323954e7 100644 --- a/Mathlib/CategoryTheory/Sites/Canonical.lean +++ b/Mathlib/CategoryTheory/Sites/Canonical.lean @@ -43,9 +43,7 @@ variable {C : Type u} [Category.{v} C] namespace Sheaf -variable {P : Cᵒᵖ ⥤ Type v} -variable {X Y : C} {S : Sieve X} {R : Presieve X} -variable (J J₂ : GrothendieckTopology C) +variable {P : Cᵒᵖ ⥤ Type v} {X : C} (J : GrothendieckTopology C) /-- To show `P` is a sheaf for the binding of `U` with `B`, it suffices to show that `P` is a sheaf for diff --git a/Mathlib/Computability/Reduce.lean b/Mathlib/Computability/Reduce.lean index f3131659de771..4b09ff9c74119 100644 --- a/Mathlib/Computability/Reduce.lean +++ b/Mathlib/Computability/Reduce.lean @@ -106,8 +106,7 @@ theorem transitive_oneOneReducible {α} [Primcodable α] : Transitive (@OneOneRe namespace ComputablePred -variable {α : Type*} {β : Type*} {σ : Type*} -variable [Primcodable α] [Primcodable β] [Primcodable σ] +variable {α : Type*} {β : Type*} [Primcodable α] [Primcodable β] open Computable @@ -266,9 +265,7 @@ theorem disjoin_le {α β γ} [Primcodable α] [Primcodable β] [Primcodable γ] OneOneReducible.disjoin_right.to_many_one.trans h⟩, fun ⟨h₁, h₂⟩ => disjoin_manyOneReducible h₁ h₂⟩ -variable {α : Type u} [Primcodable α] [Inhabited α] -variable {β : Type v} [Primcodable β] [Inhabited β] -variable {γ : Type w} [Primcodable γ] [Inhabited γ] +variable {α : Type u} [Primcodable α] [Inhabited α] {β : Type v} [Primcodable β] [Inhabited β] /-- Computable and injective mapping of predicates to sets of natural numbers. -/ diff --git a/Mathlib/Control/Fold.lean b/Mathlib/Control/Fold.lean index 71a4d437faaa3..711e71a910635 100644 --- a/Mathlib/Control/Fold.lean +++ b/Mathlib/Control/Fold.lean @@ -239,7 +239,6 @@ theorem foldl.unop_ofFreeMonoid (f : β → α → β) (xs : FreeMonoid α) (a : unop (Foldl.ofFreeMonoid f xs) a = List.foldl f a (FreeMonoid.toList xs) := rfl -variable (m : Type u → Type u) [Monad m] [LawfulMonad m] variable {t : Type u → Type u} [Traversable t] [LawfulTraversable t] open LawfulTraversable diff --git a/Mathlib/Control/Functor/Multivariate.lean b/Mathlib/Control/Functor/Multivariate.lean index 5234cb42fc0e8..042f138715b1f 100644 --- a/Mathlib/Control/Functor/Multivariate.lean +++ b/Mathlib/Control/Functor/Multivariate.lean @@ -36,7 +36,7 @@ variable {n : ℕ} namespace MvFunctor -variable {α β γ : TypeVec.{u} n} {F : TypeVec.{u} n → Type v} [MvFunctor F] +variable {α β : TypeVec.{u} n} {F : TypeVec.{u} n → Type v} [MvFunctor F] /-- predicate lifting over multivariate functors -/ def LiftP {α : TypeVec n} (P : ∀ i, α i → Prop) (x : F α) : Prop := diff --git a/Mathlib/Data/BitVec.lean b/Mathlib/Data/BitVec.lean index e3928ba4fb51c..3fb3407a84cb6 100644 --- a/Mathlib/Data/BitVec.lean +++ b/Mathlib/Data/BitVec.lean @@ -18,7 +18,7 @@ can either be PR'd to Lean, or kept downstream if it also relies on Mathlib. namespace BitVec -variable {w v : Nat} +variable {w : Nat} /-! ## Injectivity diff --git a/Mathlib/Data/DFinsupp/Multiset.lean b/Mathlib/Data/DFinsupp/Multiset.lean index 949128cc0bccb..fe589464c5396 100644 --- a/Mathlib/Data/DFinsupp/Multiset.lean +++ b/Mathlib/Data/DFinsupp/Multiset.lean @@ -14,7 +14,7 @@ with `Multiset.toDFinsupp` the reverse equivalence. open Function -variable {α : Type*} {β : α → Type*} +variable {α : Type*} namespace DFinsupp @@ -22,7 +22,7 @@ namespace DFinsupp instance addZeroClass' {β} [AddZeroClass β] : AddZeroClass (Π₀ _ : α, β) := @DFinsupp.addZeroClass α (fun _ ↦ β) _ -variable [DecidableEq α] {s t : Multiset α} +variable [DecidableEq α] /-- A DFinsupp version of `Finsupp.toMultiset`. -/ def toMultiset : (Π₀ _ : α, ℕ) →+ Multiset α := diff --git a/Mathlib/Data/TypeVec.lean b/Mathlib/Data/TypeVec.lean index ed974cab4a625..df15092ea20be 100644 --- a/Mathlib/Data/TypeVec.lean +++ b/Mathlib/Data/TypeVec.lean @@ -419,8 +419,8 @@ theorem const_iff_true {α : TypeVec n} {i x p} : ofRepeat (TypeVec.const p α i section -variable {α β γ : TypeVec.{u} n} -variable (p : α ⟹ «repeat» n Prop) (r : α ⊗ α ⟹ «repeat» n Prop) +variable {α β : TypeVec.{u} n} +variable (p : α ⟹ «repeat» n Prop) /-- left projection of a `prod` vector -/ def prod.fst : ∀ {n} {α β : TypeVec.{u} n}, α ⊗ β ⟹ α diff --git a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean index 55142a16d8ebd..d6cebfb34829f 100644 --- a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean +++ b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean @@ -417,9 +417,9 @@ abbrev ModelWithCorners.tangent {𝕜 : Type u} [NontriviallyNormedField 𝕜] { variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {E' : Type*} [NormedAddCommGroup E'] [NormedSpace 𝕜 E'] {F : Type*} - [NormedAddCommGroup F] [NormedSpace 𝕜 F] {F' : Type*} [NormedAddCommGroup F'] [NormedSpace 𝕜 F'] + [NormedAddCommGroup F] [NormedSpace 𝕜 F] {H : Type*} [TopologicalSpace H] {H' : Type*} [TopologicalSpace H'] {G : Type*} - [TopologicalSpace G] {G' : Type*} [TopologicalSpace G'] {I : ModelWithCorners 𝕜 E H} + [TopologicalSpace G] {I : ModelWithCorners 𝕜 E H} {J : ModelWithCorners 𝕜 F G} @[simp, mfld_simps] diff --git a/Mathlib/MeasureTheory/Measure/Hausdorff.lean b/Mathlib/MeasureTheory/Measure/Hausdorff.lean index fc55f1955171d..d4023bb58018d 100644 --- a/Mathlib/MeasureTheory/Measure/Hausdorff.lean +++ b/Mathlib/MeasureTheory/Measure/Hausdorff.lean @@ -669,7 +669,7 @@ variable [MeasurableSpace X] [BorelSpace X] [MeasurableSpace Y] [BorelSpace Y] namespace HolderOnWith -variable {C r : ℝ≥0} {f : X → Y} {s t : Set X} +variable {C r : ℝ≥0} {f : X → Y} {s : Set X} /-- If `f : X → Y` is Hölder continuous on `s` with a positive exponent `r`, then `μH[d] (f '' s) ≤ C ^ d * μH[r * d] s`. -/ @@ -715,7 +715,7 @@ end HolderOnWith namespace LipschitzOnWith -variable {K : ℝ≥0} {f : X → Y} {s t : Set X} +variable {K : ℝ≥0} {f : X → Y} {s : Set X} /-- If `f : X → Y` is `K`-Lipschitz on `s`, then `μH[d] (f '' s) ≤ K ^ d * μH[d] s`. -/ theorem hausdorffMeasure_image_le (h : LipschitzOnWith K f s) {d : ℝ} (hd : 0 ≤ d) : diff --git a/Mathlib/MeasureTheory/Measure/LevyProkhorovMetric.lean b/Mathlib/MeasureTheory/Measure/LevyProkhorovMetric.lean index 1a99901e0a12a..fa9209bc596e0 100644 --- a/Mathlib/MeasureTheory/Measure/LevyProkhorovMetric.lean +++ b/Mathlib/MeasureTheory/Measure/LevyProkhorovMetric.lean @@ -43,7 +43,7 @@ section Levy_Prokhorov /-! ### Lévy-Prokhorov metric -/ -variable {ι : Type*} {Ω : Type*} [MeasurableSpace Ω] [PseudoEMetricSpace Ω] +variable {Ω : Type*} [MeasurableSpace Ω] [PseudoEMetricSpace Ω] /-- The Lévy-Prokhorov edistance between measures: `d(μ,ν) = inf {r ≥ 0 | ∀ B, μ B ≤ ν Bᵣ + r ∧ ν B ≤ μ Bᵣ + r}`. -/ @@ -333,7 +333,7 @@ section Levy_Prokhorov_is_finer open BoundedContinuousFunction -variable {ι : Type*} {Ω : Type*} [MeasurableSpace Ω] +variable {Ω : Type*} [MeasurableSpace Ω] variable [PseudoMetricSpace Ω] [OpensMeasurableSpace Ω] @@ -497,7 +497,7 @@ section Levy_Prokhorov_metrizes_convergence_in_distribution open BoundedContinuousFunction TopologicalSpace -variable {ι : Type*} {Ω : Type*} [PseudoMetricSpace Ω] +variable {Ω : Type*} [PseudoMetricSpace Ω] variable [MeasurableSpace Ω] [OpensMeasurableSpace Ω] lemma ProbabilityMeasure.toMeasure_add_pos_gt_mem_nhds (P : ProbabilityMeasure Ω) diff --git a/Mathlib/MeasureTheory/Order/UpperLower.lean b/Mathlib/MeasureTheory/Order/UpperLower.lean index 6e04d112e9d58..3813db3c363dd 100644 --- a/Mathlib/MeasureTheory/Order/UpperLower.lean +++ b/Mathlib/MeasureTheory/Order/UpperLower.lean @@ -48,7 +48,7 @@ Generalize so that it also applies to `ℝ × ℝ`, for example. open Filter MeasureTheory Metric Set open scoped Topology -variable {ι : Type*} [Fintype ι] {s : Set (ι → ℝ)} {x y : ι → ℝ} {δ : ℝ} +variable {ι : Type*} [Fintype ι] {s : Set (ι → ℝ)} {x : ι → ℝ} /-- If we can fit a small ball inside a set `s` intersected with any neighborhood of `x`, then the density of `s` near `x` is not `0`. diff --git a/Mathlib/ModelTheory/Complexity.lean b/Mathlib/ModelTheory/Complexity.lean index e4d9bfc7ef31f..a44a3ddcd67f9 100644 --- a/Mathlib/ModelTheory/Complexity.lean +++ b/Mathlib/ModelTheory/Complexity.lean @@ -34,11 +34,8 @@ namespace FirstOrder namespace Language -variable {L : Language.{u, v}} {L' : Language} -variable {M : Type w} {N P : Type*} [L.Structure M] [L.Structure N] [L.Structure P] -variable {α : Type u'} {β : Type v'} {γ : Type*} -variable {n l : ℕ} {φ ψ : L.BoundedFormula α l} {θ : L.BoundedFormula α l.succ} -variable {v : α → M} {xs : Fin l → M} +variable {L : Language.{u, v}} {M : Type w} [L.Structure M] {α : Type u'} {β : Type v'} +variable {n l : ℕ} {φ : L.BoundedFormula α l} open FirstOrder Structure Fin diff --git a/Mathlib/ModelTheory/ElementarySubstructures.lean b/Mathlib/ModelTheory/ElementarySubstructures.lean index e0a575f7c8a3d..c2321b14f9d6a 100644 --- a/Mathlib/ModelTheory/ElementarySubstructures.lean +++ b/Mathlib/ModelTheory/ElementarySubstructures.lean @@ -29,8 +29,7 @@ namespace Language open Structure -variable {L : Language} {M : Type*} {N : Type*} {P : Type*} {Q : Type*} -variable [L.Structure M] [L.Structure N] [L.Structure P] [L.Structure Q] +variable {L : Language} {M : Type*} [L.Structure M] /-- A substructure is elementary when every formula applied to a tuple in the substructure agrees with its value in the overall structure. -/ diff --git a/Mathlib/ModelTheory/Encoding.lean b/Mathlib/ModelTheory/Encoding.lean index b8219e55752e1..8e4da06e5a48a 100644 --- a/Mathlib/ModelTheory/Encoding.lean +++ b/Mathlib/ModelTheory/Encoding.lean @@ -33,15 +33,14 @@ import Mathlib.SetTheory.Cardinal.Arithmetic -/ -universe u v w u' v' +universe u v w u' namespace FirstOrder namespace Language variable {L : Language.{u, v}} -variable {M : Type w} {N P : Type*} [L.Structure M] [L.Structure N] [L.Structure P] -variable {α : Type u'} {β : Type v'} +variable {α : Type u'} open FirstOrder Cardinal diff --git a/Mathlib/ModelTheory/Syntax.lean b/Mathlib/ModelTheory/Syntax.lean index 1c3ec68f882e6..24c9794e5cfbd 100644 --- a/Mathlib/ModelTheory/Syntax.lean +++ b/Mathlib/ModelTheory/Syntax.lean @@ -61,8 +61,7 @@ namespace FirstOrder namespace Language variable (L : Language.{u, v}) {L' : Language} -variable {M : Type w} {N P : Type*} [L.Structure M] [L.Structure N] [L.Structure P] -variable {α : Type u'} {β : Type v'} {γ : Type*} +variable {M : Type w} {α : Type u'} {β : Type v'} {γ : Type*} open FirstOrder diff --git a/Mathlib/RingTheory/GradedAlgebra/HomogeneousIdeal.lean b/Mathlib/RingTheory/GradedAlgebra/HomogeneousIdeal.lean index f434e98f774a3..bd67fd0c74783 100644 --- a/Mathlib/RingTheory/GradedAlgebra/HomogeneousIdeal.lean +++ b/Mathlib/RingTheory/GradedAlgebra/HomogeneousIdeal.lean @@ -44,7 +44,7 @@ open SetLike DirectSum Set open Pointwise DirectSum -variable {ι σ R A : Type*} +variable {ι σ A : Type*} section HomogeneousDef diff --git a/Mathlib/RingTheory/Localization/Finiteness.lean b/Mathlib/RingTheory/Localization/Finiteness.lean index ab340bfec9330..4998718519d38 100644 --- a/Mathlib/RingTheory/Localization/Finiteness.lean +++ b/Mathlib/RingTheory/Localization/Finiteness.lean @@ -59,7 +59,7 @@ lemma of_isLocalizedModule [Module.Finite R M] : Module.Finite Rₚ Mₚ := by end -variable {R : Type u} [CommRing R] (S : Submonoid R) {M : Type w} [AddCommMonoid M] [Module R M] +variable {R : Type u} [CommRing R] {M : Type w} [AddCommMonoid M] [Module R M] /-- If there exists a finite set `{ r }` of `R` such that `Mᵣ` is `Rᵣ`-finite for each `r`, diff --git a/Mathlib/RingTheory/TensorProduct/MvPolynomial.lean b/Mathlib/RingTheory/TensorProduct/MvPolynomial.lean index 5b8ae61882f73..529b84b2a297c 100644 --- a/Mathlib/RingTheory/TensorProduct/MvPolynomial.lean +++ b/Mathlib/RingTheory/TensorProduct/MvPolynomial.lean @@ -37,7 +37,7 @@ Let `Semiring R`, `Algebra R S` and `Module R N`. -/ -universe u v w +universe u v noncomputable section @@ -47,8 +47,7 @@ open DirectSum TensorProduct open Set LinearMap Submodule -variable {R : Type u} {M : Type v} {N : Type w} - [CommSemiring R] [AddCommMonoid M] [Module R M] +variable {R : Type u} {N : Type v} [CommSemiring R] variable {σ : Type*} diff --git a/Mathlib/Topology/MetricSpace/Holder.lean b/Mathlib/Topology/MetricSpace/Holder.lean index da0e2b14c4f81..70758a458af3f 100644 --- a/Mathlib/Topology/MetricSpace/Holder.lean +++ b/Mathlib/Topology/MetricSpace/Holder.lean @@ -239,7 +239,7 @@ end PseudoMetric section Metric -variable [PseudoMetricSpace X] [MetricSpace Y] {C r : ℝ≥0} {f : X → Y} +variable [PseudoMetricSpace X] [MetricSpace Y] {r : ℝ≥0} {f : X → Y} @[simp] lemma holderWith_zero_iff : HolderWith 0 r f ↔ ∀ x₁ x₂, f x₁ = f x₂ := by diff --git a/Mathlib/Topology/MetricSpace/MetricSeparated.lean b/Mathlib/Topology/MetricSpace/MetricSeparated.lean index 2529db5502ed7..520ee661e1219 100644 --- a/Mathlib/Topology/MetricSpace/MetricSeparated.lean +++ b/Mathlib/Topology/MetricSpace/MetricSeparated.lean @@ -27,7 +27,7 @@ def IsMetricSeparated {X : Type*} [EMetricSpace X] (s t : Set X) := namespace IsMetricSeparated -variable {X : Type*} [EMetricSpace X] {s t : Set X} {x y : X} +variable {X : Type*} [EMetricSpace X] {s t : Set X} @[symm] theorem symm (h : IsMetricSeparated s t) : IsMetricSeparated t s := diff --git a/Mathlib/Topology/Order/DenselyOrdered.lean b/Mathlib/Topology/Order/DenselyOrdered.lean index 80ad140e2daf8..e427346122b7a 100644 --- a/Mathlib/Topology/Order/DenselyOrdered.lean +++ b/Mathlib/Topology/Order/DenselyOrdered.lean @@ -13,7 +13,7 @@ open Set Filter TopologicalSpace Topology Function open OrderDual (toDual ofDual) -variable {α β γ : Type*} +variable {α β : Type*} section DenselyOrdered diff --git a/Mathlib/Topology/Order/IntermediateValue.lean b/Mathlib/Topology/Order/IntermediateValue.lean index 9b86197c50a9e..12e59d849ba91 100644 --- a/Mathlib/Topology/Order/IntermediateValue.lean +++ b/Mathlib/Topology/Order/IntermediateValue.lean @@ -45,7 +45,7 @@ intermediate value theorem, connected space, connected set open Filter OrderDual TopologicalSpace Function Set open scoped Topology Filter Interval -universe u v w +universe u v /-! ### Intermediate value theorem on a (pre)connected space @@ -216,9 +216,7 @@ theorem IsPreconnected.eq_univ_of_unbounded {s : Set α} (hs : IsPreconnected s) end -variable {α : Type u} {β : Type v} {γ : Type w} [ConditionallyCompleteLinearOrder α] - [TopologicalSpace α] [OrderTopology α] [ConditionallyCompleteLinearOrder β] [TopologicalSpace β] - [OrderTopology β] [Nonempty γ] +variable {α : Type u} [ConditionallyCompleteLinearOrder α] [TopologicalSpace α] [OrderTopology α] /-- A bounded connected subset of a conditionally complete linear order includes the open interval `(Inf s, Sup s)`. -/ diff --git a/Mathlib/Topology/Order/IsLUB.lean b/Mathlib/Topology/Order/IsLUB.lean index 2d6f0f82dc720..39f599a2851f8 100644 --- a/Mathlib/Topology/Order/IsLUB.lean +++ b/Mathlib/Topology/Order/IsLUB.lean @@ -14,12 +14,11 @@ open Set Filter TopologicalSpace Topology Function open OrderDual (toDual ofDual) -variable {α β γ : Type*} +variable {α γ : Type*} section OrderTopology -variable [TopologicalSpace α] [TopologicalSpace β] [LinearOrder α] [LinearOrder β] [OrderTopology α] - [OrderTopology β] +variable [TopologicalSpace α] [LinearOrder α] [OrderTopology α] theorem IsLUB.frequently_mem {a : α} {s : Set α} (ha : IsLUB s a) (hs : s.Nonempty) : ∃ᶠ x in 𝓝[≤] a, x ∈ s := by diff --git a/Mathlib/Topology/Order/Monotone.lean b/Mathlib/Topology/Order/Monotone.lean index c807d80f962d5..c6aefd9a4b657 100644 --- a/Mathlib/Topology/Order/Monotone.lean +++ b/Mathlib/Topology/Order/Monotone.lean @@ -19,12 +19,12 @@ open Set Filter TopologicalSpace Topology Function open OrderDual (toDual ofDual) -variable {α β γ : Type*} +variable {α β : Type*} section ConditionallyCompleteLinearOrder variable [ConditionallyCompleteLinearOrder α] [TopologicalSpace α] [OrderTopology α] - [ConditionallyCompleteLinearOrder β] [TopologicalSpace β] [OrderClosedTopology β] [Nonempty γ] + [ConditionallyCompleteLinearOrder β] [TopologicalSpace β] [OrderClosedTopology β] /-- A monotone function continuous at the supremum of a nonempty set sends this supremum to the supremum of the image of this set. -/ @@ -150,7 +150,7 @@ end ConditionallyCompleteLinearOrder section CompleteLinearOrder variable [CompleteLinearOrder α] [TopologicalSpace α] [OrderTopology α] [CompleteLinearOrder β] - [TopologicalSpace β] [OrderClosedTopology β] [Nonempty γ] + [TopologicalSpace β] [OrderClosedTopology β] theorem sSup_mem_closure {s : Set α} (hs : s.Nonempty) : sSup s ∈ closure s := (isLUB_sSup s).mem_closure hs @@ -255,7 +255,6 @@ end CompleteLinearOrder section ConditionallyCompleteLinearOrder variable [ConditionallyCompleteLinearOrder α] [TopologicalSpace α] [OrderTopology α] - [ConditionallyCompleteLinearOrder β] [TopologicalSpace β] [OrderClosedTopology β] [Nonempty γ] theorem csSup_mem_closure {s : Set α} (hs : s.Nonempty) (B : BddAbove s) : sSup s ∈ closure s := (isLUB_csSup hs B).mem_closure hs diff --git a/Mathlib/Topology/UniformSpace/Equicontinuity.lean b/Mathlib/Topology/UniformSpace/Equicontinuity.lean index 807c30865e06d..c3444a738efe4 100644 --- a/Mathlib/Topology/UniformSpace/Equicontinuity.lean +++ b/Mathlib/Topology/UniformSpace/Equicontinuity.lean @@ -82,8 +82,8 @@ section open UniformSpace Filter Set Uniformity Topology UniformConvergence Function -variable {ι κ X X' Y Z α α' β β' γ 𝓕 : Type*} [tX : TopologicalSpace X] [tY : TopologicalSpace Y] - [tZ : TopologicalSpace Z] [uα : UniformSpace α] [uβ : UniformSpace β] [uγ : UniformSpace γ] +variable {ι κ X X' Y α α' β β' γ : Type*} [tX : TopologicalSpace X] [tY : TopologicalSpace Y] + [uα : UniformSpace α] [uβ : UniformSpace β] [uγ : UniformSpace γ] /-- A family `F : ι → X → α` of functions from a topological space to a uniform space is *equicontinuous at `x₀ : X`* if, for all entourages `U ∈ 𝓤 α`, there is a neighborhood `V` of `x₀` diff --git a/Mathlib/Topology/UniformSpace/UniformConvergenceTopology.lean b/Mathlib/Topology/UniformSpace/UniformConvergenceTopology.lean index 00cd2deafccea..9a7542b8c7d0c 100644 --- a/Mathlib/Topology/UniformSpace/UniformConvergenceTopology.lean +++ b/Mathlib/Topology/UniformSpace/UniformConvergenceTopology.lean @@ -197,7 +197,7 @@ open UniformConvergence namespace UniformFun variable (α β : Type*) {γ ι : Type*} -variable {s s' : Set α} {x : α} {p : Filter ι} {g : ι → α} +variable {p : Filter ι} /-- Basis sets for the uniformity of uniform convergence: `gen α β V` is the set of pairs `(f, g)` of functions `α →ᵤ β` such that `∀ x, (f x, g x) ∈ V`. -/ @@ -555,7 +555,7 @@ end UniformFun namespace UniformOnFun variable {α β : Type*} {γ ι : Type*} -variable {s s' : Set α} {x : α} {p : Filter ι} {g : ι → α} +variable {s : Set α} {p : Filter ι} local notation "𝒰(" α ", " β ", " u ")" => @UniformFun.uniformSpace α β u From 8d6f380fde8631e39a62b5a009b2ffec2eb3c17f Mon Sep 17 00:00:00 2001 From: Yongle Hu Date: Wed, 16 Oct 2024 15:54:25 +0000 Subject: [PATCH 034/505] feat(RingTheory/Flat/Basic): drop `[Small.{v} R]` assumption in `Module.Flat.iff_*` lemmas (#17484) Drop `[Small.{v} R]` assumption from `Module.Flat.iff_*` lemmas [Module.Flat.iff_rTensor_preserves_injective_linearMap](https://leanprover-community.github.io/mathlib4_docs/Mathlib/RingTheory/Flat/Basic.html#Module.Flat.iff_rTensor_preserves_injective_linearMap), [Module.Flat.iff_lTensor_exact](https://leanprover-community.github.io/mathlib4_docs/Mathlib/RingTheory/Flat/Basic.html#Module.Flat.iff_lTensor_exact), and so on. Co-authored-by: Hu Yongle <2065545849@qq.com> --- Mathlib/RingTheory/Flat/Basic.lean | 175 +++++++++++++++++++---------- 1 file changed, 115 insertions(+), 60 deletions(-) diff --git a/Mathlib/RingTheory/Flat/Basic.lean b/Mathlib/RingTheory/Flat/Basic.lean index 813b32f84d0ac..9c3da85861bf3 100644 --- a/Mathlib/RingTheory/Flat/Basic.lean +++ b/Mathlib/RingTheory/Flat/Basic.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2020 Johan Commelin. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Johan Commelin, Jujian Zhang +Authors: Johan Commelin, Jujian Zhang, Yongle Hu -/ import Mathlib.Algebra.DirectSum.Finsupp import Mathlib.Algebra.DirectSum.Module @@ -39,14 +39,13 @@ See . * `Module.Flat.preserves_injective_linearMap`: If `M` is a flat module then tensoring with `M` preserves injectivity of linear maps. This lemma is fully universally polymorphic in all arguments, i.e. `R`, `M` and linear maps `N → N'` can all have different universe levels. -* `Module.Flat.iff_rTensor_preserves_injective_linearMap`: a module is flat iff tensoring preserves - injectivity in the ring's universe (or higher). - -## Implementation notes -In `Module.Flat.iff_rTensor_preserves_injective_linearMap`, we require that the universe level of -the ring is lower than or equal to that of the module. This requirement is to make sure ideals of -the ring can be lifted to the universe of the module. It is unclear if this lemma also holds -when the module lives in a lower universe. +* `Module.Flat.iff_rTensor_preserves_injective_linearMap`: a module is flat iff tensoring modules + in the higher universe preserves injectivity . +* `Module.Flat.lTensor_exact`: If `M` is a flat module then tensoring with `M` is an exact + functor. This lemma is fully universally polymorphic in all arguments, i.e. + `R`, `M` and linear maps `N → N' → N''` can all have different universe levels. +* `Module.Flat.iff_lTensor_exact`: a module is flat iff tensoring modules + in the higher universe is an exact functor. ## TODO @@ -55,7 +54,7 @@ when the module lives in a lower universe. -/ -universe u v w +universe v' u v w namespace Module @@ -144,6 +143,24 @@ lemma of_linearEquiv [f : Flat R M] (e : N ≃ₗ[R] M) : Flat R N := by have h : e.symm.toLinearMap.comp e.toLinearMap = LinearMap.id := by simp exact of_retract _ _ _ e.toLinearMap e.symm.toLinearMap h +/-- If an `R`-module `M` is linearly equivalent to another `R`-module `N`, then `M` is flat + if and only if `N` is flat. -/ +lemma equiv_iff (e : M ≃ₗ[R] N) : Flat R M ↔ Flat R N := + ⟨fun _ => of_linearEquiv R M N e.symm, fun _ => of_linearEquiv R N M e⟩ + +instance ulift [Module.Flat R M] : Module.Flat R (ULift.{v'} M) := + of_linearEquiv R M (ULift.{v'} M) ULift.moduleEquiv + +instance (priority := 100) of_ulift [Module.Flat R (ULift.{v'} M)] : Module.Flat R M := + of_linearEquiv R (ULift.{v'} M) M ULift.moduleEquiv.symm + +instance shrink [Small.{v'} M] [Module.Flat R M] : Module.Flat R (Shrink.{v'} M) := + of_linearEquiv R M (Shrink.{v'} M) (Shrink.linearEquiv M R) + +instance (priority := 100) of_shrink [Small.{v'} M] [Module.Flat R (Shrink.{v'} M)] : + Module.Flat R M := + of_linearEquiv R (Shrink.{v'} M) M (Shrink.linearEquiv M R).symm + /-- A direct sum of flat `R`-modules is flat. -/ instance directSum (ι : Type v) (M : ι → Type w) [(i : ι) → AddCommGroup (M i)] [(i : ι) → Module R (M i)] [F : (i : ι) → (Flat R (M i))] : Flat R (⨁ i, M i) := by @@ -250,28 +267,55 @@ theorem lTensor_preserves_injective_linearMap {N' : Type*} [AddCommGroup N'] [Mo (L.lTensor_inj_iff_rTensor_inj M).2 (rTensor_preserves_injective_linearMap L hL) variable (R M) in -/-- -M is flat if and only if `f ⊗ 𝟙 M` is injective whenever `f` is an injective linear map. --/ -lemma iff_rTensor_preserves_injective_linearMap [Small.{v} R] : - Flat R M ↔ - ∀ ⦃N N' : Type v⦄ [AddCommGroup N] [AddCommGroup N'] [Module R N] [Module R N'] - (L : N →ₗ[R] N'), Function.Injective L → Function.Injective (L.rTensor M) := by - rw [iff_characterModule_injective, - injective_characterModule_iff_rTensor_preserves_injective_linearMap] +/-- `M` is flat if and only if `f ⊗ 𝟙 M` is injective whenever `f` is an injective linear map. + See `Module.Flat.iff_rTensor_preserves_injective_linearMap` to specialize the universe of + `N, N', N''` to `Type (max u v)`. -/ +lemma iff_rTensor_preserves_injective_linearMap' [Small.{v'} R] [Small.{v'} M] : Flat R M ↔ + ∀ ⦃N N' : Type v'⦄ [AddCommGroup N] [AddCommGroup N'] [Module R N] [Module R N'] + (f : N →ₗ[R] N') (_ : Function.Injective f), Function.Injective (f.rTensor M) := + (Module.Flat.equiv_iff R M (Shrink.{v'} M) (Shrink.linearEquiv M R).symm).trans <| + iff_characterModule_injective.trans <| + (injective_characterModule_iff_rTensor_preserves_injective_linearMap R (Shrink.{v'} M)).trans + <| forall₅_congr <| fun N N' _ _ _ => forall₃_congr <| fun _ f _ => + let frmu := f.rTensor (Shrink.{v'} M) + let frm := f.rTensor M + let emn := TensorProduct.congr (LinearEquiv.refl R N) (Shrink.linearEquiv M R) + let emn' := TensorProduct.congr (LinearEquiv.refl R N') (Shrink.linearEquiv M R) + have h : emn'.toLinearMap.comp frmu = frm.comp emn.toLinearMap := TensorProduct.ext rfl + (EquivLike.comp_injective frmu emn').symm.trans <| + (congrArg Function.Injective (congrArg DFunLike.coe h)).to_iff.trans <| + EquivLike.injective_comp emn frm variable (R M) in -/-- -M is flat if and only if `𝟙 M ⊗ f` is injective whenever `f` is an injective linear map. --/ -lemma iff_lTensor_preserves_injective_linearMap [Small.{v} R] : - Flat R M ↔ - ∀ ⦃N N' : Type v⦄ [AddCommGroup N] [AddCommGroup N'] [Module R N] [Module R N'] +/-- `M` is flat if and only if `f ⊗ 𝟙 M` is injective whenever `f` is an injective linear map. + See `Module.Flat.iff_rTensor_preserves_injective_linearMap'` to generalize the universe of + `N, N', N''` to any universe that is higher than `R` and `M`. -/ +lemma iff_rTensor_preserves_injective_linearMap : Flat R M ↔ + ∀ ⦃N N' : Type (max u v)⦄ [AddCommGroup N] [AddCommGroup N'] [Module R N] [Module R N'] + (f : N →ₗ[R] N') (_ : Function.Injective f), Function.Injective (f.rTensor M) := + iff_rTensor_preserves_injective_linearMap'.{max u v} R M + +variable (R M) in +/-- `M` is flat if and only if `𝟙 M ⊗ f` is injective whenever `f` is an injective linear map. + See `Module.Flat.iff_lTensor_preserves_injective_linearMap` to specialize the universe of + `N, N', N''` to `Type (max u v)`. -/ +lemma iff_lTensor_preserves_injective_linearMap' [Small.{v'} R] [Small.{v'} M] : Flat R M ↔ + ∀ ⦃N N' : Type v'⦄ [AddCommGroup N] [AddCommGroup N'] [Module R N] [Module R N'] (L : N →ₗ[R] N'), Function.Injective L → Function.Injective (L.lTensor M) := by - simp_rw [iff_rTensor_preserves_injective_linearMap, LinearMap.lTensor_inj_iff_rTensor_inj] + simp_rw [iff_rTensor_preserves_injective_linearMap', LinearMap.lTensor_inj_iff_rTensor_inj] + +variable (R M) in +/-- `M` is flat if and only if `𝟙 M ⊗ f` is injective whenever `f` is an injective linear map. + See `Module.Flat.iff_lTensor_preserves_injective_linearMap'` to generalize the universe of + `N, N', N''` to any universe that is higher than `R` and `M`. -/ +lemma iff_lTensor_preserves_injective_linearMap : Flat R M ↔ + ∀ ⦃N N' : Type (max u v)⦄ [AddCommGroup N] [AddCommGroup N'] [Module R N] [Module R N'] + (f : N →ₗ[R] N') (_ : Function.Injective f), Function.Injective (f.lTensor M) := + iff_lTensor_preserves_injective_linearMap'.{max u v} R M variable (M) in -lemma lTensor_exact [Small.{v} R] [flat : Flat R M] ⦃N N' N'' : Type v⦄ +/-- If `M` is flat then `M ⊗ -` is an exact functor. -/ +lemma lTensor_exact [Flat R M] ⦃N N' N'' : Type*⦄ [AddCommGroup N] [AddCommGroup N'] [AddCommGroup N''] [Module R N] [Module R N'] [Module R N''] ⦃f : N →ₗ[R] N'⦄ ⦃g : N' →ₗ[R] N''⦄ (exact : Function.Exact f g) : Function.Exact (f.lTensor M) (g.lTensor M) := by @@ -280,18 +324,15 @@ lemma lTensor_exact [Small.{v} R] [flat : Flat R M] ⦃N N' N'' : Type v⦄ Submodule.subtype _ ∘ₗ (LinearMap.quotKerEquivRange g).toLinearMap ∘ₗ Submodule.quotEquivOfEq (LinearMap.range f) (LinearMap.ker g) (LinearMap.exact_iff.mp exact).symm - suffices exact1 : Function.Exact (f.lTensor M) (π.lTensor M) by - rw [show g = ι.comp π by aesop, lTensor_comp] - exact exact1.comp_injective - (inj := iff_lTensor_preserves_injective_linearMap R M |>.mp flat _ <| by - simpa [ι] using Subtype.val_injective) - (h0 := map_zero _) - + rw [show g = ι.comp π from rfl, lTensor_comp] + exact exact1.comp_injective _ (lTensor_preserves_injective_linearMap ι <| by + simpa [ι] using Subtype.val_injective) (map_zero _) exact _root_.lTensor_exact _ (fun x => by simp [π]) Quotient.surjective_Quotient_mk'' variable (M) in -lemma rTensor_exact [Small.{v} R] [flat : Flat R M] ⦃N N' N'' : Type v⦄ +/-- If `M` is flat then `- ⊗ M` is an exact functor. -/ +lemma rTensor_exact [Flat R M] ⦃N N' N'' : Type*⦄ [AddCommGroup N] [AddCommGroup N'] [AddCommGroup N''] [Module R N] [Module R N'] [Module R N''] ⦃f : N →ₗ[R] N'⦄ ⦃g : N' →ₗ[R] N''⦄ (exact : Function.Exact f g) : Function.Exact (f.rTensor M) (g.rTensor M) := by @@ -300,41 +341,55 @@ lemma rTensor_exact [Small.{v} R] [flat : Flat R M] ⦃N N' N'' : Type v⦄ Submodule.subtype _ ∘ₗ (LinearMap.quotKerEquivRange g).toLinearMap ∘ₗ Submodule.quotEquivOfEq (LinearMap.range f) (LinearMap.ker g) (LinearMap.exact_iff.mp exact).symm - suffices exact1 : Function.Exact (f.rTensor M) (π.rTensor M) by - rw [show g = ι.comp π by aesop, rTensor_comp] - exact exact1.comp_injective - (inj := iff_rTensor_preserves_injective_linearMap R M |>.mp flat _ <| by - simpa [ι] using Subtype.val_injective) - (h0 := map_zero _) - - exact _root_.rTensor_exact _ (fun x => by simp [π]) Quotient.surjective_Quotient_mk'' - -/-- -M is flat if and only if `M ⊗ -` is a left exact functor. --/ -lemma iff_lTensor_exact [Small.{v} R] : - Flat R M ↔ - ∀ ⦃N N' N'' : Type v⦄ [AddCommGroup N] [AddCommGroup N'] [AddCommGroup N''] + rw [show g = ι.comp π from rfl, rTensor_comp] + exact exact1.comp_injective _ (rTensor_preserves_injective_linearMap ι <| by + simpa [ι] using Subtype.val_injective) (map_zero _) + exact _root_.rTensor_exact M (fun x => by simp [π]) Quotient.surjective_Quotient_mk'' + +/-- `M` is flat if and only if `M ⊗ -` is an exact functor. See + `Module.Flat.iff_lTensor_exact` to specialize the universe of `N, N', N''` to `Type (max u v)`. -/ +theorem iff_lTensor_exact' [Small.{v'} R] [Small.{v'} M] : Flat R M ↔ + ∀ ⦃N N' N'' : Type v'⦄ [AddCommGroup N] [AddCommGroup N'] [AddCommGroup N''] [Module R N] [Module R N'] [Module R N''] ⦃f : N →ₗ[R] N'⦄ ⦃g : N' →ₗ[R] N''⦄, Function.Exact f g → Function.Exact (f.lTensor M) (g.lTensor M) := by - refine ⟨fun _ => lTensor_exact M, fun H => iff_lTensor_preserves_injective_linearMap R M |>.mpr + refine ⟨fun _ => lTensor_exact M, fun H => iff_lTensor_preserves_injective_linearMap' R M |>.mpr fun N' N'' _ _ _ _ L hL => LinearMap.ker_eq_bot |>.mp <| eq_bot_iff |>.mpr fun x (hx : _ = 0) => ?_⟩ - simpa [Eq.comm] using @H PUnit N' N'' _ _ _ _ _ _ 0 L (fun x => by aesop) x |>.mp hx + simpa [Eq.comm] using @H PUnit N' N'' _ _ _ _ _ _ 0 L (fun x => by + simp_rw [Set.mem_range, LinearMap.zero_apply, exists_const] + exact (L.map_eq_zero_iff hL).trans eq_comm) x |>.mp hx + +/-- `M` is flat if and only if `M ⊗ -` is an exact functor. + See `Module.Flat.iff_lTensor_exact'` to generalize the universe of + `N, N', N''` to any universe that is higher than `R` and `M`. -/ +theorem iff_lTensor_exact : Flat R M ↔ + ∀ ⦃N N' N'' : Type (max u v)⦄ [AddCommGroup N] [AddCommGroup N'] [AddCommGroup N''] + [Module R N] [Module R N'] [Module R N''] ⦃f : N →ₗ[R] N'⦄ ⦃g : N' →ₗ[R] N''⦄, + Function.Exact f g → Function.Exact (f.lTensor M) (g.lTensor M) := + iff_lTensor_exact'.{max u v} -/-- -M is flat if and only if `- ⊗ M` is a left exact functor. --/ -lemma iff_rTensor_exact [Small.{v} R] : - Flat R M ↔ - ∀ ⦃N N' N'' : Type v⦄ [AddCommGroup N] [AddCommGroup N'] [AddCommGroup N''] +/-- `M` is flat if and only if `- ⊗ M` is an exact functor. See + `Module.Flat.iff_rTensor_exact` to specialize the universe of `N, N', N''` to `Type (max u v)`. -/ +theorem iff_rTensor_exact' [Small.{v'} R] [Small.{v'} M] : Flat R M ↔ + ∀ ⦃N N' N'' : Type v'⦄ [AddCommGroup N] [AddCommGroup N'] [AddCommGroup N''] [Module R N] [Module R N'] [Module R N''] ⦃f : N →ₗ[R] N'⦄ ⦃g : N' →ₗ[R] N''⦄, Function.Exact f g → Function.Exact (f.rTensor M) (g.rTensor M) := by - refine ⟨fun _ => rTensor_exact M, fun H => iff_rTensor_preserves_injective_linearMap R M |>.mpr + refine ⟨fun _ => rTensor_exact M, fun H => iff_rTensor_preserves_injective_linearMap' R M |>.mpr fun N' N'' _ _ _ _ L hL => LinearMap.ker_eq_bot |>.mp <| eq_bot_iff |>.mpr fun x (hx : _ = 0) => ?_⟩ - simpa [Eq.comm] using @H PUnit N' N'' _ _ _ _ _ _ 0 L (fun x => by aesop) x |>.mp hx + simpa [Eq.comm] using @H PUnit N' N'' _ _ _ _ _ _ 0 L (fun x => by + simp_rw [Set.mem_range, LinearMap.zero_apply, exists_const] + exact (L.map_eq_zero_iff hL).trans eq_comm) x |>.mp hx + +/-- `M` is flat if and only if `- ⊗ M` is an exact functor. + See `Module.Flat.iff_rTensor_exact'` to generalize the universe of + `N, N', N''` to any universe that is higher than `R` and `M`. -/ +theorem iff_rTensor_exact : Flat R M ↔ + ∀ ⦃N N' N'' : Type (max u v)⦄ [AddCommGroup N] [AddCommGroup N'] [AddCommGroup N''] + [Module R N] [Module R N'] [Module R N''] ⦃f : N →ₗ[R] N'⦄ ⦃g : N' →ₗ[R] N''⦄, + Function.Exact f g → Function.Exact (f.rTensor M) (g.rTensor M) := + iff_rTensor_exact'.{max u v} end Flat From ec9e50a65ce6c7740e7676060d103be493f2169b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 16 Oct 2024 16:23:46 +0000 Subject: [PATCH 035/505] =?UTF-8?q?chore:=20`measurable=5FspanningSets`=20?= =?UTF-8?q?=E2=86=92=20`measurableSet=5FspanningSets`=20(#17834)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This was misnamed. From GibbsMeasure --- .../Constructions/BorelSpace/Real.lean | 2 +- .../Decomposition/Exhaustion.lean | 2 +- .../MeasureTheory/Decomposition/Lebesgue.lean | 4 +-- .../Function/AEEqOfIntegral.lean | 4 +-- .../ConditionalExpectation/CondexpL1.lean | 2 +- Mathlib/MeasureTheory/Function/Jacobian.lean | 4 +-- .../Function/StronglyMeasurable/Basic.lean | 2 +- Mathlib/MeasureTheory/Integral/Lebesgue.lean | 6 ++-- Mathlib/MeasureTheory/Measure/Regular.lean | 2 +- Mathlib/MeasureTheory/Measure/Trim.lean | 2 +- .../MeasureTheory/Measure/Typeclasses.lean | 30 ++++++++++--------- 11 files changed, 31 insertions(+), 29 deletions(-) diff --git a/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean b/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean index 2f7ad05237166..9dfc57bb16291 100644 --- a/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean +++ b/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean @@ -484,7 +484,7 @@ theorem exists_spanning_measurableSet_le {f : α → ℝ≥0} (hf : Measurable f let sets n := sigma_finite_sets n ∩ norm_sets n have h_meas : ∀ n, MeasurableSet (sets n) := by refine fun n => MeasurableSet.inter ?_ ?_ - · exact measurable_spanningSets μ n + · exact measurableSet_spanningSets μ n · exact hf measurableSet_Iic have h_finite : ∀ n, μ (sets n) < ∞ := by refine fun n => (measure_mono Set.inter_subset_left).trans_lt ?_ diff --git a/Mathlib/MeasureTheory/Decomposition/Exhaustion.lean b/Mathlib/MeasureTheory/Decomposition/Exhaustion.lean index a6f283377c6c5..0cf002bd37fb1 100644 --- a/Mathlib/MeasureTheory/Decomposition/Exhaustion.lean +++ b/Mathlib/MeasureTheory/Decomposition/Exhaustion.lean @@ -284,7 +284,7 @@ lemma measure_compl_sigmaFiniteSetWRT (hμν : μ ≪ ν) [SigmaFinite μ] [SFin by_contra h0 refine ENNReal.top_ne_zero ?_ rw [← h h0, ← Measure.iSup_restrict_spanningSets] - simp_rw [Measure.restrict_apply' (measurable_spanningSets μ _), ENNReal.iSup_eq_zero] + simp_rw [Measure.restrict_apply' (measurableSet_spanningSets μ _), ENNReal.iSup_eq_zero] intro i by_contra h_ne_zero have h_zero_top := measure_eq_top_of_subset_compl_sigmaFiniteSetWRT diff --git a/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean b/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean index b82573cc25bc5..291ea3c448fc5 100644 --- a/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean +++ b/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean @@ -369,7 +369,7 @@ theorem rnDeriv_lt_top (μ ν : Measure α) [SigmaFinite μ] : ∀ᵐ x ∂ν, suffices ∀ n, ∀ᵐ x ∂ν, x ∈ spanningSets μ n → μ.rnDeriv ν x < ∞ by filter_upwards [ae_all_iff.2 this] with _ hx using hx _ (mem_spanningSetsIndex _ _) intro n - rw [← ae_restrict_iff' (measurable_spanningSets _ _)] + rw [← ae_restrict_iff' (measurableSet_spanningSets _ _)] apply ae_lt_top (measurable_rnDeriv _ _) refine (lintegral_rnDeriv_lt_top_of_measure_ne_top _ ?_).ne exact (measure_spanningSets_lt_top _ _).ne @@ -938,7 +938,7 @@ nonrec instance (priority := 100) haveLebesgueDecomposition_of_sigmaFinite · exact .sfinite_of_isFiniteMeasure fun μ _ ↦ this μ ‹_› -- Take a disjoint cover that consists of sets of finite measure `ν`. set s : ℕ → Set α := disjointed (spanningSets ν) - have hsm : ∀ n, MeasurableSet (s n) := .disjointed <| measurable_spanningSets _ + have hsm : ∀ n, MeasurableSet (s n) := .disjointed <| measurableSet_spanningSets _ have hs : ∀ n, Fact (ν (s n) < ⊤) := fun n ↦ ⟨lt_of_le_of_lt (measure_mono <| disjointed_le ..) (measure_spanningSets_lt_top ν n)⟩ -- Note that the restrictions of `μ` and `ν` to `s n` are finite measures. diff --git a/Mathlib/MeasureTheory/Function/AEEqOfIntegral.lean b/Mathlib/MeasureTheory/Function/AEEqOfIntegral.lean index d9303573bf9fd..311f212e8f62f 100644 --- a/Mathlib/MeasureTheory/Function/AEEqOfIntegral.lean +++ b/Mathlib/MeasureTheory/Function/AEEqOfIntegral.lean @@ -423,10 +423,10 @@ theorem ae_eq_zero_of_forall_setIntegral_eq_of_sigmaFinite [SigmaFinite μ] {f : (hf_zero : ∀ s : Set α, MeasurableSet s → μ s < ∞ → ∫ x in s, f x ∂μ = 0) : f =ᵐ[μ] 0 := by let S := spanningSets μ rw [← @Measure.restrict_univ _ _ μ, ← iUnion_spanningSets μ, EventuallyEq, ae_iff, - Measure.restrict_apply' (MeasurableSet.iUnion (measurable_spanningSets μ))] + Measure.restrict_apply' (MeasurableSet.iUnion (measurableSet_spanningSets μ))] rw [Set.inter_iUnion, measure_iUnion_null_iff] intro n - have h_meas_n : MeasurableSet (S n) := measurable_spanningSets μ n + have h_meas_n : MeasurableSet (S n) := measurableSet_spanningSets μ n have hμn : μ (S n) < ∞ := measure_spanningSets_lt_top μ n rw [← Measure.restrict_apply' h_meas_n] exact ae_eq_zero_restrict_of_forall_setIntegral_eq_zero hf_int_finite hf_zero h_meas_n hμn.ne diff --git a/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL1.lean b/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL1.lean index a0b70b96c1884..633ff287c5581 100644 --- a/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL1.lean +++ b/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL1.lean @@ -395,7 +395,7 @@ to the integral of `f` on that set. See also `setIntegral_condexp`, the similar theorem setIntegral_condexpL1CLM (f : α →₁[μ] F') (hs : MeasurableSet[m] s) : ∫ x in s, condexpL1CLM F' hm μ f x ∂μ = ∫ x in s, f x ∂μ := by let S := spanningSets (μ.trim hm) - have hS_meas : ∀ i, MeasurableSet[m] (S i) := measurable_spanningSets (μ.trim hm) + have hS_meas : ∀ i, MeasurableSet[m] (S i) := measurableSet_spanningSets (μ.trim hm) have hS_meas0 : ∀ i, MeasurableSet (S i) := fun i => hm _ (hS_meas i) have hs_eq : s = ⋃ i, S i ∩ s := by simp_rw [Set.inter_comm] diff --git a/Mathlib/MeasureTheory/Function/Jacobian.lean b/Mathlib/MeasureTheory/Function/Jacobian.lean index ce6116b747cb5..77a1878659c77 100644 --- a/Mathlib/MeasureTheory/Function/Jacobian.lean +++ b/Mathlib/MeasureTheory/Function/Jacobian.lean @@ -883,7 +883,7 @@ theorem addHaar_image_le_lintegral_abs_det_fderiv (hs : MeasurableSet s) have u_meas : ∀ n, MeasurableSet (u n) := by intro n apply MeasurableSet.disjointed fun i => ?_ - exact measurable_spanningSets μ i + exact measurableSet_spanningSets μ i have A : s = ⋃ n, s ∩ u n := by rw [← inter_iUnion, iUnion_disjointed, iUnion_spanningSets, inter_univ] calc @@ -1035,7 +1035,7 @@ theorem lintegral_abs_det_fderiv_le_addHaar_image (hs : MeasurableSet s) have u_meas : ∀ n, MeasurableSet (u n) := by intro n apply MeasurableSet.disjointed fun i => ?_ - exact measurable_spanningSets μ i + exact measurableSet_spanningSets μ i have A : s = ⋃ n, s ∩ u n := by rw [← inter_iUnion, iUnion_disjointed, iUnion_spanningSets, inter_univ] calc diff --git a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean index b9644215bcca2..7d3943e021a7c 100644 --- a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean +++ b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean @@ -269,7 +269,7 @@ theorem finStronglyMeasurable_of_set_sigmaFinite [TopologicalSpace β] [Zero β] FinStronglyMeasurable f μ := by haveI : SigmaFinite (μ.restrict t) := htμ let S := spanningSets (μ.restrict t) - have hS_meas : ∀ n, MeasurableSet (S n) := measurable_spanningSets (μ.restrict t) + have hS_meas : ∀ n, MeasurableSet (S n) := measurableSet_spanningSets (μ.restrict t) let f_approx := hf_meas.approx let fs n := SimpleFunc.restrict (f_approx n) (S n ∩ t) have h_fs_t_compl : ∀ n, ∀ x, x ∉ t → fs n x = 0 := by diff --git a/Mathlib/MeasureTheory/Integral/Lebesgue.lean b/Mathlib/MeasureTheory/Integral/Lebesgue.lean index ea728df43b140..5be3b2d1f299b 100644 --- a/Mathlib/MeasureTheory/Integral/Lebesgue.lean +++ b/Mathlib/MeasureTheory/Integral/Lebesgue.lean @@ -1898,11 +1898,11 @@ theorem exists_pos_lintegral_lt_of_sigmaFinite (μ : Measure α) [SigmaFinite μ obtain ⟨δ, δpos, δsum⟩ : ∃ δ : ℕ → ℝ≥0, (∀ i, 0 < δ i) ∧ (∑' i, μ (s i) * δ i) < ε := ENNReal.exists_pos_tsum_mul_lt_of_countable ε0 _ fun n => (this n).ne set N : α → ℕ := spanningSetsIndex μ - have hN_meas : Measurable N := measurable_spanningSetsIndex μ + have hN_meas : Measurable N := measurableSet_spanningSetsIndex μ have hNs : ∀ n, N ⁻¹' {n} = s n := preimage_spanningSetsIndex_singleton μ refine ⟨δ ∘ N, fun x => δpos _, measurable_from_nat.comp hN_meas, ?_⟩ erw [lintegral_comp measurable_from_nat.coe_nnreal_ennreal hN_meas] - simpa [N, hNs, lintegral_countable', measurable_spanningSetsIndex, mul_comm] using δsum + simpa [N, hNs, lintegral_countable', measurableSet_spanningSetsIndex, mul_comm] using δsum theorem lintegral_trim {μ : Measure α} (hm : m ≤ m0) {f : α → ℝ≥0∞} (hf : Measurable[m] f) : ∫⁻ a, f a ∂μ.trim hm = ∫⁻ a, f a ∂μ := by @@ -1939,7 +1939,7 @@ theorem univ_le_of_forall_fin_meas_le {μ : Measure α} (hm : m ≤ m0) [SigmaFi f univ ≤ C := by let S := @spanningSets _ m (μ.trim hm) _ have hS_mono : Monotone S := @monotone_spanningSets _ m (μ.trim hm) _ - have hS_meas : ∀ n, MeasurableSet[m] (S n) := @measurable_spanningSets _ m (μ.trim hm) _ + have hS_meas : ∀ n, MeasurableSet[m] (S n) := @measurableSet_spanningSets _ m (μ.trim hm) _ rw [← @iUnion_spanningSets _ m (μ.trim hm)] refine (h_F_lim S hS_meas hS_mono).trans ?_ refine iSup_le fun n => hf (S n) (hS_meas n) ?_ diff --git a/Mathlib/MeasureTheory/Measure/Regular.lean b/Mathlib/MeasureTheory/Measure/Regular.lean index 362b5a971e2a4..f283ccf0b3e83 100644 --- a/Mathlib/MeasureTheory/Measure/Regular.lean +++ b/Mathlib/MeasureTheory/Measure/Regular.lean @@ -511,7 +511,7 @@ lemma of_sigmaFinite [SigmaFinite μ] : rw [← (monotone_const.inter (monotone_spanningSets μ)).measure_iUnion, hBU] rw [this] at hr rcases lt_iSup_iff.1 hr with ⟨n, hn⟩ - refine ⟨s ∩ B n, inter_subset_left, ⟨hs.inter (measurable_spanningSets μ n), ?_⟩, hn⟩ + refine ⟨s ∩ B n, inter_subset_left, ⟨hs.inter (measurableSet_spanningSets μ n), ?_⟩, hn⟩ exact ((measure_mono inter_subset_right).trans_lt (measure_spanningSets_lt_top μ n)).ne variable [TopologicalSpace α] diff --git a/Mathlib/MeasureTheory/Measure/Trim.lean b/Mathlib/MeasureTheory/Measure/Trim.lean index af9a03aa6916a..3afc9ad874998 100644 --- a/Mathlib/MeasureTheory/Measure/Trim.lean +++ b/Mathlib/MeasureTheory/Measure/Trim.lean @@ -104,7 +104,7 @@ theorem sigmaFiniteTrim_mono {m m₂ m0 : MeasurableSpace α} {μ : Measure α} calc (μ.trim hm) (spanningSets (μ.trim (hm₂.trans hm)) i) = ((μ.trim hm).trim hm₂) (spanningSets (μ.trim (hm₂.trans hm)) i) := by - rw [@trim_measurableSet_eq α m₂ m (μ.trim hm) _ hm₂ (measurable_spanningSets _ _)] + rw [@trim_measurableSet_eq α m₂ m (μ.trim hm) _ hm₂ (measurableSet_spanningSets _ _)] _ = (μ.trim (hm₂.trans hm)) (spanningSets (μ.trim (hm₂.trans hm)) i) := by rw [@trim_trim _ _ μ _ _ hm₂ hm] _ < ∞ := measure_spanningSets_lt_top _ _ diff --git a/Mathlib/MeasureTheory/Measure/Typeclasses.lean b/Mathlib/MeasureTheory/Measure/Typeclasses.lean index 5dc07c2bfba72..888fe3d0777ee 100644 --- a/Mathlib/MeasureTheory/Measure/Typeclasses.lean +++ b/Mathlib/MeasureTheory/Measure/Typeclasses.lean @@ -672,10 +672,12 @@ theorem monotone_spanningSets (μ : Measure α) [SigmaFinite μ] : Monotone (spa lemma spanningSets_mono [SigmaFinite μ] {m n : ℕ} (hmn : m ≤ n) : spanningSets μ m ⊆ spanningSets μ n := monotone_spanningSets _ hmn -theorem measurable_spanningSets (μ : Measure α) [SigmaFinite μ] (i : ℕ) : +theorem measurableSet_spanningSets (μ : Measure α) [SigmaFinite μ] (i : ℕ) : MeasurableSet (spanningSets μ i) := MeasurableSet.iUnion fun j => MeasurableSet.iUnion fun _ => μ.toFiniteSpanningSetsIn.set_mem j +@[deprecated (since := "2024-10-16")] alias measurable_spanningSets := measurableSet_spanningSets + theorem measure_spanningSets_lt_top (μ : Measure α) [SigmaFinite μ] (i : ℕ) : μ (spanningSets μ i) < ∞ := measure_biUnion_lt_top (finite_le_nat i) fun j _ => μ.toFiniteSpanningSetsIn.finite j @@ -694,9 +696,9 @@ noncomputable def spanningSetsIndex (μ : Measure α) [SigmaFinite μ] (x : α) Nat.find <| iUnion_eq_univ_iff.1 (iUnion_spanningSets μ) x open scoped Classical in -theorem measurable_spanningSetsIndex (μ : Measure α) [SigmaFinite μ] : +theorem measurableSet_spanningSetsIndex (μ : Measure α) [SigmaFinite μ] : Measurable (spanningSetsIndex μ) := - measurable_find _ <| measurable_spanningSets μ + measurable_find _ <| measurableSet_spanningSets μ open scoped Classical in theorem preimage_spanningSetsIndex_singleton (μ : Measure α) [SigmaFinite μ] (n : ℕ) : @@ -726,7 +728,7 @@ theorem eventually_mem_spanningSets (μ : Measure α) [SigmaFinite μ] (x : α) theorem sum_restrict_disjointed_spanningSets (μ ν : Measure α) [SigmaFinite ν] : sum (fun n ↦ μ.restrict (disjointed (spanningSets ν) n)) = μ := by rw [← restrict_iUnion (disjoint_disjointed _) - (MeasurableSet.disjointed (measurable_spanningSets _)), + (MeasurableSet.disjointed (measurableSet_spanningSets _)), iUnion_disjointed, iUnion_spanningSets, restrict_univ] instance (priority := 100) [SigmaFinite μ] : SFinite μ := by @@ -988,9 +990,9 @@ theorem iSup_restrict_spanningSets [SigmaFinite μ] (s : Set α) : ⨆ i, μ.restrict (spanningSets μ i) s = μ s := by rw [← measure_toMeasurable s, ← iSup_restrict_spanningSets_of_measurableSet (measurableSet_toMeasurable _ _)] - simp_rw [restrict_apply' (measurable_spanningSets μ _), Set.inter_comm s, - ← restrict_apply (measurable_spanningSets μ _), ← restrict_toMeasurable_of_sFinite s, - restrict_apply (measurable_spanningSets μ _), Set.inter_comm _ (toMeasurable μ s)] + simp_rw [restrict_apply' (measurableSet_spanningSets μ _), Set.inter_comm s, + ← restrict_apply (measurableSet_spanningSets μ _), ← restrict_toMeasurable_of_sFinite s, + restrict_apply (measurableSet_spanningSets μ _), Set.inter_comm _ (toMeasurable μ s)] /-- In a σ-finite space, any measurable set of measure `> r` contains a measurable subset of finite measure `> r`. -/ @@ -1001,7 +1003,7 @@ theorem exists_subset_measure_lt_top [SigmaFinite μ] {r : ℝ≥0∞} (hs : Mea rcases h's with ⟨n, hn⟩ simp only [restrict_apply hs] at hn refine - ⟨s ∩ spanningSets μ n, hs.inter (measurable_spanningSets _ _), inter_subset_left, hn, ?_⟩ + ⟨s ∩ spanningSets μ n, hs.inter (measurableSet_spanningSets _ _), inter_subset_left, hn, ?_⟩ exact (measure_mono inter_subset_right).trans_lt (measure_spanningSets_lt_top _ _) namespace FiniteSpanningSetsIn @@ -1081,7 +1083,7 @@ theorem sigmaFinite_bot_iff (μ : @Measure α ⊥) : SigmaFinite μ ↔ IsFinite haveI : SigmaFinite μ := h let s := spanningSets μ have hs_univ : ⋃ i, s i = Set.univ := iUnion_spanningSets μ - have hs_meas : ∀ i, MeasurableSet[⊥] (s i) := measurable_spanningSets μ + have hs_meas : ∀ i, MeasurableSet[⊥] (s i) := measurableSet_spanningSets μ simp_rw [MeasurableSpace.measurableSet_bot_iff] at hs_meas by_cases h_univ_empty : (Set.univ : Set α) = ∅ · rw [h_univ_empty, measure_empty] @@ -1097,14 +1099,14 @@ theorem sigmaFinite_bot_iff (μ : @Measure α ⊥) : SigmaFinite μ ↔ IsFinite instance Restrict.sigmaFinite (μ : Measure α) [SigmaFinite μ] (s : Set α) : SigmaFinite (μ.restrict s) := by refine ⟨⟨⟨spanningSets μ, fun _ => trivial, fun i => ?_, iUnion_spanningSets μ⟩⟩⟩ - rw [Measure.restrict_apply (measurable_spanningSets μ i)] + rw [Measure.restrict_apply (measurableSet_spanningSets μ i)] exact (measure_mono inter_subset_left).trans_lt (measure_spanningSets_lt_top μ i) instance sum.sigmaFinite {ι} [Finite ι] (μ : ι → Measure α) [∀ i, SigmaFinite (μ i)] : SigmaFinite (sum μ) := by cases nonempty_fintype ι have : ∀ n, MeasurableSet (⋂ i : ι, spanningSets (μ i) n) := fun n => - MeasurableSet.iInter fun i => measurable_spanningSets (μ i) n + MeasurableSet.iInter fun i => measurableSet_spanningSets (μ i) n refine ⟨⟨⟨fun n => ⋂ i, spanningSets (μ i) n, fun _ => trivial, fun n => ?_, ?_⟩⟩⟩ · rw [sum_apply _ (this n), tsum_fintype, ENNReal.sum_lt_top] rintro i - @@ -1141,7 +1143,7 @@ instance [SigmaFinite (μ.restrict t)] : SigmaFinite (μ.restrict (s ∩ t)) := theorem SigmaFinite.of_map (μ : Measure α) {f : α → β} (hf : AEMeasurable f μ) (h : SigmaFinite (μ.map f)) : SigmaFinite μ := ⟨⟨⟨fun n => f ⁻¹' spanningSets (μ.map f) n, fun _ => trivial, fun n => by - simp only [← map_apply_of_aemeasurable hf, measurable_spanningSets, + simp only [← map_apply_of_aemeasurable hf, measurableSet_spanningSets, measure_spanningSets_lt_top], by rw [← preimage_iUnion, iUnion_spanningSets, preimage_univ]⟩⟩⟩ @@ -1167,10 +1169,10 @@ theorem ae_of_forall_measure_lt_top_ae_restrict' {μ : Measure α} (ν : Measure have : ∀ n, ∀ᵐ x ∂μ, x ∈ spanningSets (μ + ν) n → P x := by intro n have := h - (spanningSets (μ + ν) n) (measurable_spanningSets _ _) + (spanningSets (μ + ν) n) (measurableSet_spanningSets _ _) ((self_le_add_right _ _).trans_lt (measure_spanningSets_lt_top (μ + ν) _)) ((self_le_add_left _ _).trans_lt (measure_spanningSets_lt_top (μ + ν) _)) - exact (ae_restrict_iff' (measurable_spanningSets _ _)).mp this + exact (ae_restrict_iff' (measurableSet_spanningSets _ _)).mp this filter_upwards [ae_all_iff.2 this] with _ hx using hx _ (mem_spanningSetsIndex _ _) /-- To prove something for almost all `x` w.r.t. a σ-finite measure, it is sufficient to show that From 2de78712a1ca26b255be2351ff54b919d3cba7ba Mon Sep 17 00:00:00 2001 From: damiano Date: Wed, 16 Oct 2024 17:21:07 +0000 Subject: [PATCH 036/505] chore: remove more unused variables (#17749) --- Mathlib/Algebra/Group/Action/Pi.lean | 2 +- Mathlib/Algebra/Group/Units/Defs.lean | 6 +++--- Mathlib/Algebra/GroupPower/IterateHom.lean | 2 +- Mathlib/Algebra/Order/Group/Prod.lean | 2 -- Mathlib/Algebra/Order/Monovary.lean | 12 ++++++------ .../Algebra/Order/Ring/Unbundled/Nonneg.lean | 7 ------- Mathlib/Algebra/Order/Ring/Unbundled/Rat.lean | 2 +- Mathlib/Algebra/Ring/Hom/Basic.lean | 4 ++-- Mathlib/Algebra/Ring/Parity.lean | 4 ++-- .../BoxIntegral/Partition/Additive.lean | 2 +- Mathlib/Analysis/Calculus/Deriv/Basic.lean | 3 +-- Mathlib/Analysis/Calculus/Deriv/Comp.lean | 12 ++++++------ Mathlib/Analysis/Calculus/Deriv/Inv.lean | 17 +++++------------ Mathlib/Analysis/Calculus/Deriv/Polynomial.lean | 16 ++++------------ Mathlib/Analysis/Calculus/Deriv/Pow.lean | 16 ++-------------- Mathlib/Analysis/Calculus/FDeriv/Bilinear.lean | 10 +--------- Mathlib/Analysis/Calculus/FDeriv/Linear.lean | 16 ++++------------ Mathlib/Analysis/Convex/EGauge.lean | 14 ++++++-------- Mathlib/Analysis/Normed/Operator/Banach.lean | 2 +- .../Analysis/NormedSpace/OperatorNorm/Mul.lean | 2 +- .../PiTensorProduct/ProjectiveSeminorm.lean | 1 - .../Functor/KanExtension/Basic.lean | 3 +-- .../CategoryTheory/Limits/Shapes/Images.lean | 4 ++-- Mathlib/CategoryTheory/Limits/Types.lean | 2 +- Mathlib/Condensed/Light/CartesianClosed.lean | 2 -- Mathlib/Data/Fintype/Option.lean | 2 +- Mathlib/Data/NNRat/Defs.lean | 4 ++-- Mathlib/Data/Set/Basic.lean | 4 ++-- Mathlib/Data/Set/Image.lean | 2 +- Mathlib/Data/Set/Prod.lean | 2 +- Mathlib/Dynamics/Ergodic/Conservative.lean | 2 +- Mathlib/Dynamics/Ergodic/Ergodic.lean | 2 +- Mathlib/Geometry/Manifold/ContMDiff/Defs.lean | 15 --------------- .../Geometry/Manifold/ContMDiff/Product.lean | 8 +------- Mathlib/Geometry/Manifold/MFDeriv/Tangent.lean | 4 ---- Mathlib/LinearAlgebra/DFinsupp.lean | 17 ++++++----------- Mathlib/MeasureTheory/Decomposition/Jordan.lean | 4 ++-- Mathlib/MeasureTheory/Function/Egorov.lean | 4 ++-- .../Integral/LebesgueNormedSpace.lean | 2 +- Mathlib/MeasureTheory/Measure/Complex.lean | 2 +- .../MeasureTheory/Measure/VectorMeasure.lean | 2 +- Mathlib/Order/UpperLower/Basic.lean | 6 +++--- .../Disintegration/MeasurableStieltjes.lean | 2 +- .../ProbabilityMassFunction/Basic.lean | 9 ++++----- .../ContinuousMap/StoneWeierstrass.lean | 2 +- Mathlib/Topology/MetricSpace/Pseudo/Basic.lean | 9 ++++----- 46 files changed, 88 insertions(+), 179 deletions(-) diff --git a/Mathlib/Algebra/Group/Action/Pi.lean b/Mathlib/Algebra/Group/Action/Pi.lean index d8c0c7b05deb3..f1ed601c36220 100644 --- a/Mathlib/Algebra/Group/Action/Pi.lean +++ b/Mathlib/Algebra/Group/Action/Pi.lean @@ -21,7 +21,7 @@ This file defines instances for `MulAction` and related structures on `Pi` types assert_not_exists MonoidWithZero -variable {ι M N : Type*} {α β γ : ι → Type*} (x y : ∀ i, α i) (i : ι) +variable {ι M N : Type*} {α β γ : ι → Type*} (i : ι) namespace Pi diff --git a/Mathlib/Algebra/Group/Units/Defs.lean b/Mathlib/Algebra/Group/Units/Defs.lean index 3e67d4529c21e..f603439477532 100644 --- a/Mathlib/Algebra/Group/Units/Defs.lean +++ b/Mathlib/Algebra/Group/Units/Defs.lean @@ -170,7 +170,7 @@ the `AddMonoid`."] instance [Repr α] : Repr αˣ := ⟨reprPrec ∘ val⟩ -variable (a b c : αˣ) {u : αˣ} +variable (a b : αˣ) {u : αˣ} @[to_additive (attr := simp, norm_cast)] theorem val_mul : (↑(a * b) : α) = a * b := @@ -477,7 +477,7 @@ theorem mul_iff [CommMonoid M] {x y : M} : IsUnit (x * y) ↔ IsUnit x ∧ IsUni section Monoid -variable [Monoid M] {a b c : M} +variable [Monoid M] {a : M} /-- The element of the group of units, corresponding to an element of a monoid which is a unit. When `α` is a `DivisionMonoid`, use `IsUnit.unit'` instead. -/ @@ -577,7 +577,7 @@ protected lemma mul_div_mul_right (h : IsUnit c) (a b : α) : a * c / (b * c) = end DivisionMonoid section DivisionCommMonoid -variable [DivisionCommMonoid α] {a b c d : α} +variable [DivisionCommMonoid α] {a c : α} @[to_additive] protected lemma div_mul_cancel_left (h : IsUnit a) (b : α) : a / (a * b) = b⁻¹ := by diff --git a/Mathlib/Algebra/GroupPower/IterateHom.lean b/Mathlib/Algebra/GroupPower/IterateHom.lean index d28361dadaa6e..50ac0d58a372a 100644 --- a/Mathlib/Algebra/GroupPower/IterateHom.lean +++ b/Mathlib/Algebra/GroupPower/IterateHom.lean @@ -31,7 +31,7 @@ assert_not_exists DenselyOrdered open Function -variable {M : Type*} {N : Type*} {G : Type*} {H : Type*} +variable {M : Type*} {G : Type*} {H : Type*} /-- An auxiliary lemma that can be used to prove `⇑(f ^ n) = ⇑f^[n]`. -/ theorem hom_coe_pow {F : Type*} [Monoid F] (c : F → M → M) (h1 : c 1 = id) diff --git a/Mathlib/Algebra/Order/Group/Prod.lean b/Mathlib/Algebra/Order/Group/Prod.lean index 3d325521fa712..9faaea7bf1205 100644 --- a/Mathlib/Algebra/Order/Group/Prod.lean +++ b/Mathlib/Algebra/Order/Group/Prod.lean @@ -11,8 +11,6 @@ import Mathlib.Algebra.Order.Monoid.Prod -/ -variable {α : Type*} - namespace Prod variable {G H : Type*} diff --git a/Mathlib/Algebra/Order/Monovary.lean b/Mathlib/Algebra/Order/Monovary.lean index e085b565b3140..60f025df1e810 100644 --- a/Mathlib/Algebra/Order/Monovary.lean +++ b/Mathlib/Algebra/Order/Monovary.lean @@ -26,7 +26,7 @@ variable {ι α β : Type*} /-! ### Algebraic operations on monovarying functions -/ section OrderedCommGroup -variable [OrderedCommGroup α] [OrderedCommGroup β] {s : Set ι} {f f₁ f₂ : ι → α} {g g₁ g₂ : ι → β} +variable [OrderedCommGroup α] [OrderedCommGroup β] {s : Set ι} {f f₁ f₂ : ι → α} {g : ι → β} @[to_additive (attr := simp)] lemma monovaryOn_inv_left : MonovaryOn f⁻¹ g s ↔ AntivaryOn f g s := by @@ -118,7 +118,7 @@ lemma Antivary.div_left (h₁ : Antivary f₁ g) (h₂ : Monovary f₂ g) : Anti end OrderedCommGroup section LinearOrderedCommGroup -variable [OrderedCommGroup α] [LinearOrderedCommGroup β] {s : Set ι} {f f₁ f₂ : ι → α} +variable [OrderedCommGroup α] [LinearOrderedCommGroup β] {s : Set ι} {f : ι → α} {g g₁ g₂ : ι → β} @[to_additive] lemma MonovaryOn.mul_right (h₁ : MonovaryOn f g₁ s) (h₂ : MonovaryOn f g₂ s) : @@ -168,7 +168,7 @@ variable [OrderedCommGroup α] [LinearOrderedCommGroup β] {s : Set ι} {f f₁ end LinearOrderedCommGroup section OrderedSemiring -variable [OrderedSemiring α] [OrderedSemiring β] {s : Set ι} {f f₁ f₂ : ι → α} {g g₁ g₂ : ι → β} +variable [OrderedSemiring α] [OrderedSemiring β] {s : Set ι} {f f₁ f₂ : ι → α} {g : ι → β} lemma MonovaryOn.mul_left₀ (hf₁ : ∀ i ∈ s, 0 ≤ f₁ i) (hf₂ : ∀ i ∈ s, 0 ≤ f₂ i) (h₁ : MonovaryOn f₁ g s) (h₂ : MonovaryOn f₂ g s) : MonovaryOn (f₁ * f₂) g s := @@ -201,7 +201,7 @@ lemma Antivary.pow_left₀ (hf : 0 ≤ f) (hfg : Antivary f g) (n : ℕ) : Antiv end OrderedSemiring section LinearOrderedSemiring -variable [LinearOrderedSemiring α] [LinearOrderedSemiring β] {s : Set ι} {f f₁ f₂ : ι → α} +variable [LinearOrderedSemiring α] [LinearOrderedSemiring β] {s : Set ι} {f : ι → α} {g g₁ g₂ : ι → β} lemma MonovaryOn.mul_right₀ (hg₁ : ∀ i ∈ s, 0 ≤ g₁ i) (hg₂ : ∀ i ∈ s, 0 ≤ g₂ i) @@ -325,7 +325,7 @@ end LinearOrderedSemifield section LinearOrderedAddCommGroup variable [LinearOrderedRing α] [LinearOrderedAddCommGroup β] [Module α β] - [OrderedSMul α β] {f : ι → α} {g : ι → β} {s : Set ι} {a a₁ a₂ : α} {b b₁ b₂ : β} + [OrderedSMul α β] {f : ι → α} {g : ι → β} {s : Set ι} lemma monovaryOn_iff_forall_smul_nonneg : MonovaryOn f g s ↔ ∀ ⦃i⦄, i ∈ s → ∀ ⦃j⦄, j ∈ s → 0 ≤ (f j - f i) • (g j - g i) := by @@ -380,7 +380,7 @@ alias ⟨AntivaryOn.smul_add_smul_le_smul_add_smul, _⟩ := antivaryOn_iff_smul_ end LinearOrderedAddCommGroup section LinearOrderedRing -variable [LinearOrderedRing α] {f g : ι → α} {s : Set ι} {a b c d : α} +variable [LinearOrderedRing α] {f g : ι → α} {s : Set ι} lemma monovaryOn_iff_forall_mul_nonneg : MonovaryOn f g s ↔ ∀ ⦃i⦄, i ∈ s → ∀ ⦃j⦄, j ∈ s → 0 ≤ (f j - f i) * (g j - g i) := by diff --git a/Mathlib/Algebra/Order/Ring/Unbundled/Nonneg.lean b/Mathlib/Algebra/Order/Ring/Unbundled/Nonneg.lean index af0b158eae60c..672f46dd53d38 100644 --- a/Mathlib/Algebra/Order/Ring/Unbundled/Nonneg.lean +++ b/Mathlib/Algebra/Order/Ring/Unbundled/Nonneg.lean @@ -273,13 +273,6 @@ def coeRingHom : { x : α // 0 ≤ x } →+* α := end Semiring -section Nontrivial - -variable [Semiring α] [PartialOrder α] [ZeroLEOneClass α] [NeZero 1] - [CovariantClass α α (· + ·) (· ≤ ·)] [PosMulMono α] - -end Nontrivial - section CommSemiring variable [CommSemiring α] [PartialOrder α] [ZeroLEOneClass α] diff --git a/Mathlib/Algebra/Order/Ring/Unbundled/Rat.lean b/Mathlib/Algebra/Order/Ring/Unbundled/Rat.lean index b2ad2a602b376..857fea4b6cfde 100644 --- a/Mathlib/Algebra/Order/Ring/Unbundled/Rat.lean +++ b/Mathlib/Algebra/Order/Ring/Unbundled/Rat.lean @@ -29,7 +29,7 @@ assert_not_exists GaloisConnection namespace Rat -variable {a b c p q : ℚ} +variable {a b p q : ℚ} @[simp] lemma divInt_nonneg_iff_of_pos_right {a b : ℤ} (hb : 0 < b) : 0 ≤ a /. b ↔ 0 ≤ a := by cases' hab : a /. b with n d hd hnd diff --git a/Mathlib/Algebra/Ring/Hom/Basic.lean b/Mathlib/Algebra/Ring/Hom/Basic.lean index 8fd264112dd85..354e0701e8d83 100644 --- a/Mathlib/Algebra/Ring/Hom/Basic.lean +++ b/Mathlib/Algebra/Ring/Hom/Basic.lean @@ -21,13 +21,13 @@ They could be moved to more natural homes. open Function -variable {F α β γ : Type*} +variable {α β : Type*} namespace RingHom section -variable {_ : NonAssocSemiring α} {_ : NonAssocSemiring β} (f : α →+* β) {x y : α} +variable {_ : NonAssocSemiring α} {_ : NonAssocSemiring β} (f : α →+* β) /-- `f : α →+* β` has a trivial codomain iff its range is `{0}`. -/ theorem codomain_trivial_iff_range_eq_singleton_zero : (0 : β) = 1 ↔ Set.range f = {0} := diff --git a/Mathlib/Algebra/Ring/Parity.lean b/Mathlib/Algebra/Ring/Parity.lean index 9205b9aeff65a..6adab468d1261 100644 --- a/Mathlib/Algebra/Ring/Parity.lean +++ b/Mathlib/Algebra/Ring/Parity.lean @@ -31,7 +31,7 @@ assert_not_exists OrderedRing open MulOpposite -variable {F α β R : Type*} +variable {F α β : Type*} section Monoid variable [Monoid α] [HasDistribNeg α] {n : ℕ} {a : α} @@ -156,7 +156,7 @@ lemma Odd.pow_add_pow_eq_zero [IsCancelAdd α] (hn : Odd n) (hab : a + b = 0) : end Semiring section Monoid -variable [Monoid α] [HasDistribNeg α] {a : α} {n : ℕ} +variable [Monoid α] [HasDistribNeg α] {n : ℕ} lemma Odd.neg_pow : Odd n → ∀ a : α, (-a) ^ n = -a ^ n := by rintro ⟨c, rfl⟩ a; simp_rw [pow_add, pow_mul, neg_sq, pow_one, mul_neg] diff --git a/Mathlib/Analysis/BoxIntegral/Partition/Additive.lean b/Mathlib/Analysis/BoxIntegral/Partition/Additive.lean index a82e54fd54bb5..648e9dfc503a9 100644 --- a/Mathlib/Analysis/BoxIntegral/Partition/Additive.lean +++ b/Mathlib/Analysis/BoxIntegral/Partition/Additive.lean @@ -57,7 +57,7 @@ namespace BoxAdditiveMap open Box Prepartition Finset -variable {N : Type*} [AddCommMonoid M] [AddCommMonoid N] {I₀ : WithTop (Box ι)} {I J : Box ι} +variable {N : Type*} [AddCommMonoid M] [AddCommMonoid N] {I₀ : WithTop (Box ι)} {I : Box ι} {i : ι} instance : FunLike (ι →ᵇᵃ[I₀] M) (Box ι) M where diff --git a/Mathlib/Analysis/Calculus/Deriv/Basic.lean b/Mathlib/Analysis/Calculus/Deriv/Basic.lean index 5ed888191d0e0..91f89a76a1777 100644 --- a/Mathlib/Analysis/Calculus/Deriv/Basic.lean +++ b/Mathlib/Analysis/Calculus/Deriv/Basic.lean @@ -97,7 +97,6 @@ open ContinuousLinearMap (smulRight smulRight_one_eq_iff) variable {𝕜 : Type u} [NontriviallyNormedField 𝕜] variable {F : Type v} [NormedAddCommGroup F] [NormedSpace 𝕜 F] -variable {E : Type w} [NormedAddCommGroup E] [NormedSpace 𝕜 E] /-- `f` has the derivative `f'` at the point `x` as `x` goes along the filter `L`. @@ -142,7 +141,7 @@ If the derivative exists (i.e., `∃ f', HasDerivAt f f' x`), then def deriv (f : 𝕜 → F) (x : 𝕜) := fderiv 𝕜 f x 1 -variable {f f₀ f₁ g : 𝕜 → F} +variable {f f₀ f₁ : 𝕜 → F} variable {f' f₀' f₁' g' : F} variable {x : 𝕜} variable {s t : Set 𝕜} diff --git a/Mathlib/Analysis/Calculus/Deriv/Comp.lean b/Mathlib/Analysis/Calculus/Deriv/Comp.lean index bb0e4bdc37b52..b57c7cc2a7dcf 100644 --- a/Mathlib/Analysis/Calculus/Deriv/Comp.lean +++ b/Mathlib/Analysis/Calculus/Deriv/Comp.lean @@ -43,11 +43,11 @@ open ContinuousLinearMap (smulRight smulRight_one_eq_iff) variable {𝕜 : Type u} [NontriviallyNormedField 𝕜] variable {F : Type v} [NormedAddCommGroup F] [NormedSpace 𝕜 F] variable {E : Type w} [NormedAddCommGroup E] [NormedSpace 𝕜 E] -variable {f f₀ f₁ g : 𝕜 → F} -variable {f' f₀' f₁' g' : F} +variable {f : 𝕜 → F} +variable {f' : F} variable {x : 𝕜} -variable {s t : Set 𝕜} -variable {L L₁ L₂ : Filter 𝕜} +variable {s : Set 𝕜} +variable {L : Filter 𝕜} section Composition @@ -65,8 +65,8 @@ usual multiplication in `comp` lemmas. /- For composition lemmas, we put x explicit to help the elaborator, as otherwise Lean tends to get confused since there are too many possibilities for composition -/ variable {𝕜' : Type*} [NontriviallyNormedField 𝕜'] [NormedAlgebra 𝕜 𝕜'] [NormedSpace 𝕜' F] - [IsScalarTower 𝕜 𝕜' F] {s' t' : Set 𝕜'} {h : 𝕜 → 𝕜'} {h₁ : 𝕜 → 𝕜} {h₂ : 𝕜' → 𝕜'} {h' h₂' : 𝕜'} - {h₁' : 𝕜} {g₁ : 𝕜' → F} {g₁' : F} {L' : Filter 𝕜'} {y : 𝕜'} (x) + [IsScalarTower 𝕜 𝕜' F] {s' t' : Set 𝕜'} {h : 𝕜 → 𝕜'} {h₂ : 𝕜' → 𝕜'} {h' h₂' : 𝕜'} + {g₁ : 𝕜' → F} {g₁' : F} {L' : Filter 𝕜'} {y : 𝕜'} (x) theorem HasDerivAtFilter.scomp (hg : HasDerivAtFilter g₁ g₁' (h x) L') (hh : HasDerivAtFilter h h' x L) (hL : Tendsto h L L') : diff --git a/Mathlib/Analysis/Calculus/Deriv/Inv.lean b/Mathlib/Analysis/Calculus/Deriv/Inv.lean index 8687b3af9373d..045ec40200bc6 100644 --- a/Mathlib/Analysis/Calculus/Deriv/Inv.lean +++ b/Mathlib/Analysis/Calculus/Deriv/Inv.lean @@ -21,21 +21,14 @@ derivative -/ -universe u v w +universe u -open scoped Classical Topology ENNReal +open scoped Topology open Filter Asymptotics Set -open ContinuousLinearMap (smulRight smulRight_one_eq_iff) +open ContinuousLinearMap (smulRight) -variable {𝕜 : Type u} [NontriviallyNormedField 𝕜] -variable {F : Type v} [NormedAddCommGroup F] [NormedSpace 𝕜 F] -variable {E : Type w} [NormedAddCommGroup E] [NormedSpace 𝕜 E] -variable {f f₀ f₁ g : 𝕜 → F} -variable {f' f₀' f₁' g' : F} -variable {x : 𝕜} -variable {s t : Set 𝕜} -variable {L : Filter 𝕜} +variable {𝕜 : Type u} [NontriviallyNormedField 𝕜] {x : 𝕜} {s : Set 𝕜} section Inverse @@ -101,7 +94,7 @@ theorem fderivWithin_inv (x_ne_zero : x ≠ 0) (hxs : UniqueDiffWithinAt 𝕜 s rw [DifferentiableAt.fderivWithin (differentiableAt_inv x_ne_zero) hxs] exact fderiv_inv -variable {c : 𝕜 → 𝕜} {h : E → 𝕜} {c' : 𝕜} {z : E} {S : Set E} +variable {c : 𝕜 → 𝕜} {c' : 𝕜} theorem HasDerivWithinAt.inv (hc : HasDerivWithinAt c c' s x) (hx : c x ≠ 0) : HasDerivWithinAt (fun y => (c y)⁻¹) (-c' / c x ^ 2) s x := by diff --git a/Mathlib/Analysis/Calculus/Deriv/Polynomial.lean b/Mathlib/Analysis/Calculus/Deriv/Polynomial.lean index 3174575fb194d..1e4623488005d 100644 --- a/Mathlib/Analysis/Calculus/Deriv/Polynomial.lean +++ b/Mathlib/Analysis/Calculus/Deriv/Polynomial.lean @@ -28,21 +28,13 @@ derivative, polynomial -/ -universe u v w +universe u -open scoped Topology Filter ENNReal Polynomial -open Set +open scoped Polynomial -open ContinuousLinearMap (smulRight smulRight_one_eq_iff) +open ContinuousLinearMap (smulRight) -variable {𝕜 : Type u} [NontriviallyNormedField 𝕜] -variable {F : Type v} [NormedAddCommGroup F] [NormedSpace 𝕜 F] -variable {E : Type w} [NormedAddCommGroup E] [NormedSpace 𝕜 E] -variable {f f₀ f₁ g : 𝕜 → F} -variable {f' f₀' f₁' g' : F} -variable {x : 𝕜} -variable {s t : Set 𝕜} -variable {L L₁ L₂ : Filter 𝕜} +variable {𝕜 : Type u} [NontriviallyNormedField 𝕜] {x : 𝕜} {s : Set 𝕜} namespace Polynomial diff --git a/Mathlib/Analysis/Calculus/Deriv/Pow.lean b/Mathlib/Analysis/Calculus/Deriv/Pow.lean index 91052a01837fb..07f1c232737a5 100644 --- a/Mathlib/Analysis/Calculus/Deriv/Pow.lean +++ b/Mathlib/Analysis/Calculus/Deriv/Pow.lean @@ -19,21 +19,9 @@ For a more detailed overview of one-dimensional derivatives in mathlib, see the derivative, power -/ -universe u v w +universe u -open scoped Classical -open Topology Filter ENNReal - -open Filter Asymptotics Set - -variable {𝕜 : Type u} [NontriviallyNormedField 𝕜] -variable {F : Type v} [NormedAddCommGroup F] [NormedSpace 𝕜 F] -variable {E : Type w} [NormedAddCommGroup E] [NormedSpace 𝕜 E] -variable {f f₀ f₁ g : 𝕜 → F} -variable {f' f₀' f₁' g' : F} -variable {x : 𝕜} -variable {s t : Set 𝕜} -variable {L L₁ L₂ : Filter 𝕜} +variable {𝕜 : Type u} [NontriviallyNormedField 𝕜] {x : 𝕜} {s : Set 𝕜} /-! ### Derivative of `x ↦ x^n` for `n : ℕ` -/ diff --git a/Mathlib/Analysis/Calculus/FDeriv/Bilinear.lean b/Mathlib/Analysis/Calculus/FDeriv/Bilinear.lean index b5afecc0aa536..b16d9f55939ba 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Bilinear.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Bilinear.lean @@ -16,9 +16,7 @@ bounded bilinear maps. -/ -open Filter Asymptotics ContinuousLinearMap Set Metric -open scoped Classical -open Topology NNReal Asymptotics ENNReal +open Asymptotics Topology noncomputable section @@ -29,12 +27,6 @@ variable {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] variable {F : Type*} [NormedAddCommGroup F] [NormedSpace 𝕜 F] variable {G : Type*} [NormedAddCommGroup G] [NormedSpace 𝕜 G] variable {G' : Type*} [NormedAddCommGroup G'] [NormedSpace 𝕜 G'] -variable {f f₀ f₁ g : E → F} -variable {f' f₀' f₁' g' : E →L[𝕜] F} -variable (e : E →L[𝕜] F) -variable {x : E} -variable {s t : Set E} -variable {L L₁ L₂ : Filter E} section BilinearMap diff --git a/Mathlib/Analysis/Calculus/FDeriv/Linear.lean b/Mathlib/Analysis/Calculus/FDeriv/Linear.lean index eba65ff473261..a705e49b74306 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Linear.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Linear.lean @@ -17,26 +17,18 @@ bounded linear maps. -/ -open Filter Asymptotics ContinuousLinearMap Set Metric - -open scoped Classical -open Topology NNReal Filter Asymptotics ENNReal - -noncomputable section +open Asymptotics section variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] variable {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] variable {F : Type*} [NormedAddCommGroup F] [NormedSpace 𝕜 F] -variable {G : Type*} [NormedAddCommGroup G] [NormedSpace 𝕜 G] -variable {G' : Type*} [NormedAddCommGroup G'] [NormedSpace 𝕜 G'] -variable {f f₀ f₁ g : E → F} -variable {f' f₀' f₁' g' : E →L[𝕜] F} +variable {f : E → F} variable (e : E →L[𝕜] F) variable {x : E} -variable {s t : Set E} -variable {L L₁ L₂ : Filter E} +variable {s : Set E} +variable {L : Filter E} section ContinuousLinearMap diff --git a/Mathlib/Analysis/Convex/EGauge.lean b/Mathlib/Analysis/Convex/EGauge.lean index 0fc1a2b1524c1..dacb554ee7505 100644 --- a/Mathlib/Analysis/Convex/EGauge.lean +++ b/Mathlib/Analysis/Convex/EGauge.lean @@ -65,8 +65,7 @@ end SMul section SMulZero -variable (𝕜 : Type*) [NNNorm 𝕜] [Nonempty 𝕜] {E : Type*} [Zero E] [SMulZeroClass 𝕜 E] - {c : 𝕜} {s t : Set E} {x : E} {r : ℝ≥0∞} +variable (𝕜 : Type*) [NNNorm 𝕜] [Nonempty 𝕜] {E : Type*} [Zero E] [SMulZeroClass 𝕜 E] {x : E} @[simp] lemma egauge_zero_left_eq_top : egauge 𝕜 0 x = ∞ ↔ x ≠ 0 := by simp [egauge_eq_top] @@ -77,8 +76,8 @@ end SMulZero section Module -variable {𝕜 : Type*} [NormedDivisionRing 𝕜] {α E : Type*} [AddCommGroup E] [Module 𝕜 E] - {c : 𝕜} {s t : Set E} {x y : E} {r : ℝ≥0∞} +variable {𝕜 : Type*} [NormedDivisionRing 𝕜] {E : Type*} [AddCommGroup E] [Module 𝕜 E] + {c : 𝕜} {s : Set E} {x : E} /-- If `c • x ∈ s` and `c ≠ 0`, then `egauge 𝕜 s x` is at most `((‖c‖₊⁻¹ : ℝ≥0) : ℝ≥0∞). @@ -159,8 +158,7 @@ end Module section SeminormedAddCommGroup -variable (𝕜 : Type*) [NormedField 𝕜] {α E : Type*} - [SeminormedAddCommGroup E] [NormedSpace 𝕜 E] {c : 𝕜} {s t : Set E} {x y : E} +variable (𝕜 : Type*) [NormedField 𝕜] {E : Type*} [SeminormedAddCommGroup E] [NormedSpace 𝕜 E] lemma div_le_egauge_closedBall (r : ℝ≥0) (x : E) : ‖x‖₊ / r ≤ egauge 𝕜 (closedBall 0 r) x := by rw [le_egauge_iff] @@ -183,8 +181,8 @@ end SeminormedAddCommGroup section SeminormedAddCommGroup -variable {𝕜 : Type*} [NormedField 𝕜] {α E : Type*} - [NormedAddCommGroup E] [NormedSpace 𝕜 E] {c : 𝕜} {s t : Set E} {x y : E} {r : ℝ≥0} +variable {𝕜 : Type*} [NormedField 𝕜] {E : Type*} + [NormedAddCommGroup E] [NormedSpace 𝕜 E] {c : 𝕜} {x : E} {r : ℝ≥0} lemma egauge_ball_le_of_one_lt_norm (hc : 1 < ‖c‖) (h₀ : r ≠ 0 ∨ x ≠ 0) : egauge 𝕜 (ball 0 r) x ≤ ‖c‖₊ * ‖x‖₊ / r := by diff --git a/Mathlib/Analysis/Normed/Operator/Banach.lean b/Mathlib/Analysis/Normed/Operator/Banach.lean index 8561d451ba771..33ad7e49abced 100644 --- a/Mathlib/Analysis/Normed/Operator/Banach.lean +++ b/Mathlib/Analysis/Normed/Operator/Banach.lean @@ -531,7 +531,7 @@ section BijectivityCriteria namespace ContinuousLinearMap -variable {σ : 𝕜 →+* 𝕜'} {σ' : 𝕜' →+* 𝕜} [RingHomInvPair σ σ'] {f : E →SL[σ] F} +variable {σ : 𝕜 →+* 𝕜'} {σ' : 𝕜' →+* 𝕜} [RingHomInvPair σ σ'] variable {F : Type u_4} [NormedAddCommGroup F] [NormedSpace 𝕜' F] variable [CompleteSpace E] diff --git a/Mathlib/Analysis/NormedSpace/OperatorNorm/Mul.lean b/Mathlib/Analysis/NormedSpace/OperatorNorm/Mul.lean index a42244f918fa7..25796540eece9 100644 --- a/Mathlib/Analysis/NormedSpace/OperatorNorm/Mul.lean +++ b/Mathlib/Analysis/NormedSpace/OperatorNorm/Mul.lean @@ -239,7 +239,7 @@ section Normed namespace ContinuousLinearMap -variable [NormedAddCommGroup E] [NormedSpace 𝕜 E] (c : 𝕜) +variable [NormedAddCommGroup E] [NormedSpace 𝕜 E] variable (𝕜) (𝕜' : Type*) section diff --git a/Mathlib/Analysis/NormedSpace/PiTensorProduct/ProjectiveSeminorm.lean b/Mathlib/Analysis/NormedSpace/PiTensorProduct/ProjectiveSeminorm.lean index ec8cbeca4457f..edb1a33efca0f 100644 --- a/Mathlib/Analysis/NormedSpace/PiTensorProduct/ProjectiveSeminorm.lean +++ b/Mathlib/Analysis/NormedSpace/PiTensorProduct/ProjectiveSeminorm.lean @@ -40,7 +40,6 @@ universe uι u𝕜 uE uF variable {ι : Type uι} [Fintype ι] variable {𝕜 : Type u𝕜} [NontriviallyNormedField 𝕜] variable {E : ι → Type uE} [∀ i, SeminormedAddCommGroup (E i)] -variable {F : Type uF} [SeminormedAddCommGroup F] [NormedSpace 𝕜 F] open scoped TensorProduct diff --git a/Mathlib/CategoryTheory/Functor/KanExtension/Basic.lean b/Mathlib/CategoryTheory/Functor/KanExtension/Basic.lean index 32996f6d67fcb..6b966a9ee45ed 100644 --- a/Mathlib/CategoryTheory/Functor/KanExtension/Basic.lean +++ b/Mathlib/CategoryTheory/Functor/KanExtension/Basic.lean @@ -36,8 +36,7 @@ open Category Limits namespace Functor -variable {C C' H H' D D' : Type*} [Category C] [Category C'] [Category H] [Category H'] - [Category D] [Category D'] +variable {C C' H D D' : Type*} [Category C] [Category C'] [Category H] [Category D] [Category D'] /-- Given two functors `L : C ⥤ D` and `F : C ⥤ H`, this is the category of functors `F' : H ⥤ D` equipped with a natural transformation `L ⋙ F' ⟶ F`. -/ diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Images.lean b/Mathlib/CategoryTheory/Limits/Shapes/Images.lean index 10e2d68009de0..fec3f30cad1ff 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Images.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Images.lean @@ -756,7 +756,7 @@ end HasImageMap section -variable (C) [Category.{v} C] [HasImages C] +variable (C) [HasImages C] /-- If a category `has_image_maps`, then all commutative squares induce morphisms on images. -/ class HasImageMaps : Prop where @@ -830,7 +830,7 @@ end StrongEpiMonoFactorisation section HasStrongEpiImages -variable (C) [Category.{v} C] [HasImages C] +variable (C) [HasImages C] /-- A category has strong epi images if it has all images and `factorThruImage f` is a strong epimorphism for all `f`. -/ diff --git a/Mathlib/CategoryTheory/Limits/Types.lean b/Mathlib/CategoryTheory/Limits/Types.lean index b67691c5eca11..e278c77547635 100644 --- a/Mathlib/CategoryTheory/Limits/Types.lean +++ b/Mathlib/CategoryTheory/Limits/Types.lean @@ -634,7 +634,7 @@ instance : HasImageMaps (Type u) where simp only [Functor.id_obj, Functor.id_map, types_comp_apply] at p rw [p, Classical.choose_spec x.2]⟩⟩) rfl -variable {F : ℕᵒᵖ ⥤ Type u} {c : Cone F} (hc : IsLimit c) +variable {F : ℕᵒᵖ ⥤ Type u} {c : Cone F} (hF : ∀ n, Function.Surjective (F.map (homOfLE (Nat.le_succ n)).op)) private noncomputable def limitOfSurjectionsSurjective.preimage diff --git a/Mathlib/Condensed/Light/CartesianClosed.lean b/Mathlib/Condensed/Light/CartesianClosed.lean index 99af9b717d59a..4d443905053a2 100644 --- a/Mathlib/Condensed/Light/CartesianClosed.lean +++ b/Mathlib/Condensed/Light/CartesianClosed.lean @@ -17,6 +17,4 @@ noncomputable section open CategoryTheory -variable {C : Type u} [SmallCategory C] - instance : CartesianClosed (LightCondSet.{u}) := inferInstanceAs (CartesianClosed (Sheaf _ _)) diff --git a/Mathlib/Data/Fintype/Option.lean b/Mathlib/Data/Fintype/Option.lean index 1a4f57cb48b3e..5bb9fafa8f916 100644 --- a/Mathlib/Data/Fintype/Option.lean +++ b/Mathlib/Data/Fintype/Option.lean @@ -19,7 +19,7 @@ open Nat universe u v -variable {α β γ : Type*} +variable {α β : Type*} open Finset Function diff --git a/Mathlib/Data/NNRat/Defs.lean b/Mathlib/Data/NNRat/Defs.lean index 00cb00c61e506..5bda9ba5b3826 100644 --- a/Mathlib/Data/NNRat/Defs.lean +++ b/Mathlib/Data/NNRat/Defs.lean @@ -61,7 +61,7 @@ deriving instance Inhabited for NNRat namespace NNRat -variable {α : Type*} {p q : ℚ≥0} +variable {p q : ℚ≥0} @[simp] lemma val_eq_cast (q : ℚ≥0) : q.1 = q := rfl @@ -345,7 +345,7 @@ See also `Rat.divInt` and `mkRat`. -/ def divNat (n d : ℕ) : ℚ≥0 := ⟨.divInt n d, Rat.divInt_nonneg (Int.ofNat_zero_le n) (Int.ofNat_zero_le d)⟩ -variable {n₁ n₂ d₁ d₂ d : ℕ} +variable {n₁ n₂ d₁ d₂ : ℕ} @[simp, norm_cast] lemma coe_divNat (n d : ℕ) : (divNat n d : ℚ) = .divInt n d := rfl diff --git a/Mathlib/Data/Set/Basic.lean b/Mathlib/Data/Set/Basic.lean index 691836860b538..2cc7a256bcfdf 100644 --- a/Mathlib/Data/Set/Basic.lean +++ b/Mathlib/Data/Set/Basic.lean @@ -65,7 +65,7 @@ set, sets, subset, subsets, union, intersection, insert, singleton, complement, open Function -universe u v w x +universe u v namespace Set @@ -182,7 +182,7 @@ theorem Eq.subset {α} {s t : Set α} : s = t → s ⊆ t := namespace Set -variable {α : Type u} {β : Type v} {γ : Type w} {ι : Sort x} {a b : α} {s s₁ s₂ t t₁ t₂ u : Set α} +variable {α : Type u} {β : Type v} {a b : α} {s s₁ s₂ t t₁ t₂ u : Set α} instance : Inhabited (Set α) := ⟨∅⟩ diff --git a/Mathlib/Data/Set/Image.lean b/Mathlib/Data/Set/Image.lean index bbd73befb9a8c..b0d95b24f9381 100644 --- a/Mathlib/Data/Set/Image.lean +++ b/Mathlib/Data/Set/Image.lean @@ -37,7 +37,7 @@ open Function Set namespace Set -variable {α β γ : Type*} {ι ι' : Sort*} +variable {α β γ : Type*} {ι : Sort*} /-! ### Inverse image -/ diff --git a/Mathlib/Data/Set/Prod.lean b/Mathlib/Data/Set/Prod.lean index 7b460fdb9cb95..5464761ef4475 100644 --- a/Mathlib/Data/Set/Prod.lean +++ b/Mathlib/Data/Set/Prod.lean @@ -532,7 +532,7 @@ namespace Set section OffDiag -variable {α : Type*} {s t : Set α} {x : α × α} {a : α} +variable {α : Type*} {s t : Set α} {a : α} theorem offDiag_mono : Monotone (offDiag : Set α → Set (α × α)) := fun _ _ h _ => And.imp (@h _) <| And.imp_left <| @h _ diff --git a/Mathlib/Dynamics/Ergodic/Conservative.lean b/Mathlib/Dynamics/Ergodic/Conservative.lean index f89c15dd99f1b..2b9a33bb435ba 100644 --- a/Mathlib/Dynamics/Ergodic/Conservative.lean +++ b/Mathlib/Dynamics/Ergodic/Conservative.lean @@ -40,7 +40,7 @@ noncomputable section open Set Filter MeasureTheory Finset Function TopologicalSpace Topology -variable {ι : Type*} {α : Type*} [MeasurableSpace α] {f : α → α} {s : Set α} {μ : Measure α} +variable {α : Type*} [MeasurableSpace α] {f : α → α} {s : Set α} {μ : Measure α} namespace MeasureTheory diff --git a/Mathlib/Dynamics/Ergodic/Ergodic.lean b/Mathlib/Dynamics/Ergodic/Ergodic.lean index eb2fbac4d06d4..6b1648fa292e1 100644 --- a/Mathlib/Dynamics/Ergodic/Ergodic.lean +++ b/Mathlib/Dynamics/Ergodic/Ergodic.lean @@ -80,7 +80,7 @@ end PreErgodic namespace MeasureTheory.MeasurePreserving -variable {β : Type*} {m' : MeasurableSpace β} {μ' : Measure β} {s' : Set β} {g : α → β} +variable {β : Type*} {m' : MeasurableSpace β} {μ' : Measure β} {g : α → β} theorem preErgodic_of_preErgodic_conjugate (hg : MeasurePreserving g μ μ') (hf : PreErgodic f μ) {f' : β → β} (h_comm : Semiconj g f f') : PreErgodic f' μ' where diff --git a/Mathlib/Geometry/Manifold/ContMDiff/Defs.lean b/Mathlib/Geometry/Manifold/ContMDiff/Defs.lean index cde9ee6782c47..e52bf1aa65520 100644 --- a/Mathlib/Geometry/Manifold/ContMDiff/Defs.lean +++ b/Mathlib/Geometry/Manifold/ContMDiff/Defs.lean @@ -64,21 +64,6 @@ variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E'' : Type*} [NormedAddCommGroup E''] [NormedSpace 𝕜 E''] {H'' : Type*} [TopologicalSpace H''] {I'' : ModelWithCorners 𝕜 E'' H''} {M'' : Type*} [TopologicalSpace M''] [ChartedSpace H'' M''] - -- declare a smooth manifold `N` over the pair `(F, G)`. - {F : Type*} - [NormedAddCommGroup F] [NormedSpace 𝕜 F] {G : Type*} [TopologicalSpace G] - {J : ModelWithCorners 𝕜 F G} {N : Type*} [TopologicalSpace N] [ChartedSpace G N] - [SmoothManifoldWithCorners J N] - -- declare a smooth manifold `N'` over the pair `(F', G')`. - {F' : Type*} - [NormedAddCommGroup F'] [NormedSpace 𝕜 F'] {G' : Type*} [TopologicalSpace G'] - {J' : ModelWithCorners 𝕜 F' G'} {N' : Type*} [TopologicalSpace N'] [ChartedSpace G' N'] - [SmoothManifoldWithCorners J' N'] - -- F₁, F₂, F₃, F₄ are normed spaces - {F₁ : Type*} - [NormedAddCommGroup F₁] [NormedSpace 𝕜 F₁] {F₂ : Type*} [NormedAddCommGroup F₂] - [NormedSpace 𝕜 F₂] {F₃ : Type*} [NormedAddCommGroup F₃] [NormedSpace 𝕜 F₃] {F₄ : Type*} - [NormedAddCommGroup F₄] [NormedSpace 𝕜 F₄] -- declare functions, sets, points and smoothness indices {e : PartialHomeomorph M H} {e' : PartialHomeomorph M' H'} {f f₁ : M → M'} {s s₁ t : Set M} {x : M} {m n : ℕ∞} diff --git a/Mathlib/Geometry/Manifold/ContMDiff/Product.lean b/Mathlib/Geometry/Manifold/ContMDiff/Product.lean index 5208fce1bdcc3..5c76e05789139 100644 --- a/Mathlib/Geometry/Manifold/ContMDiff/Product.lean +++ b/Mathlib/Geometry/Manifold/ContMDiff/Product.lean @@ -36,14 +36,8 @@ variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {F' : Type*} [NormedAddCommGroup F'] [NormedSpace 𝕜 F'] {G' : Type*} [TopologicalSpace G'] {J' : ModelWithCorners 𝕜 F' G'} {N' : Type*} [TopologicalSpace N'] [ChartedSpace G' N'] - -- F₁, F₂, F₃, F₄ are normed spaces - {F₁ : Type*} - [NormedAddCommGroup F₁] [NormedSpace 𝕜 F₁] {F₂ : Type*} [NormedAddCommGroup F₂] - [NormedSpace 𝕜 F₂] {F₃ : Type*} [NormedAddCommGroup F₃] [NormedSpace 𝕜 F₃] {F₄ : Type*} - [NormedAddCommGroup F₄] [NormedSpace 𝕜 F₄] -- declare functions, sets, points and smoothness indices - {e : PartialHomeomorph M H} - {e' : PartialHomeomorph M' H'} {f f₁ : M → M'} {s s₁ t : Set M} {x : M} {m n : ℕ∞} + {f : M → M'} {s : Set M} {x : M} {n : ℕ∞} variable {I I'} section ProdMk diff --git a/Mathlib/Geometry/Manifold/MFDeriv/Tangent.lean b/Mathlib/Geometry/Manifold/MFDeriv/Tangent.lean index b1da80bc215d6..633857d4ac0f7 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/Tangent.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/Tangent.lean @@ -20,10 +20,6 @@ open Bundle variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] (I : ModelWithCorners 𝕜 E H) {M : Type*} [TopologicalSpace M] [ChartedSpace H M] - {E' : Type*} [NormedAddCommGroup E'] [NormedSpace 𝕜 E'] {H' : Type*} [TopologicalSpace H'] - (I' : ModelWithCorners 𝕜 E' H') {M' : Type*} [TopologicalSpace M'] [ChartedSpace H' M'] - {E'' : Type*} [NormedAddCommGroup E''] [NormedSpace 𝕜 E''] {H'' : Type*} [TopologicalSpace H''] - (I'' : ModelWithCorners 𝕜 E'' H'') {M'' : Type*} [TopologicalSpace M''] [ChartedSpace H'' M''] [SmoothManifoldWithCorners I M] /-- The derivative of the chart at a base point is the chart of the tangent bundle, composed with diff --git a/Mathlib/LinearAlgebra/DFinsupp.lean b/Mathlib/LinearAlgebra/DFinsupp.lean index 2313a713d4ae5..292f5f03db793 100644 --- a/Mathlib/LinearAlgebra/DFinsupp.lean +++ b/Mathlib/LinearAlgebra/DFinsupp.lean @@ -531,18 +531,13 @@ namespace LinearMap section AddCommMonoid -variable {R : Type*} {R₁ : Type*} {R₂ : Type*} {R₃ : Type*} {R₄ : Type*} -variable {S : Type*} -variable {K : Type*} {K₂ : Type*} -variable {M : Type*} {M' : Type*} {M₁ : Type*} {M₂ : Type*} {M₃ : Type*} {M₄ : Type*} -variable {N : Type*} {N₂ : Type*} +variable {R : Type*} {R₂ : Type*} +variable {M : Type*} {M₂ : Type*} variable {ι : Type*} -variable {V : Type*} {V₂ : Type*} -variable [Semiring R] [Semiring R₂] [Semiring R₃] -variable [AddCommMonoid M] [AddCommMonoid M₂] [AddCommMonoid M₃] -variable {σ₁₂ : R →+* R₂} {σ₂₃ : R₂ →+* R₃} {σ₁₃ : R →+* R₃} -variable [RingHomCompTriple σ₁₂ σ₂₃ σ₁₃] -variable [Module R M] [Module R₂ M₂] [Module R₃ M₃] +variable [Semiring R] [Semiring R₂] +variable [AddCommMonoid M] [AddCommMonoid M₂] +variable {σ₁₂ : R →+* R₂} +variable [Module R M] [Module R₂ M₂] open Submodule diff --git a/Mathlib/MeasureTheory/Decomposition/Jordan.lean b/Mathlib/MeasureTheory/Decomposition/Jordan.lean index 6ff1e74d9497e..12402fc826cce 100644 --- a/Mathlib/MeasureTheory/Decomposition/Jordan.lean +++ b/Mathlib/MeasureTheory/Decomposition/Jordan.lean @@ -44,7 +44,7 @@ noncomputable section open scoped Classical MeasureTheory ENNReal NNReal -variable {α β : Type*} [MeasurableSpace α] +variable {α : Type*} [MeasurableSpace α] namespace MeasureTheory @@ -187,7 +187,7 @@ namespace SignedMeasure open scoped Classical open JordanDecomposition Measure Set VectorMeasure -variable {s : SignedMeasure α} {μ ν : Measure α} [IsFiniteMeasure μ] [IsFiniteMeasure ν] +variable {s : SignedMeasure α} /-- Given a signed measure `s`, `s.toJordanDecomposition` is the Jordan decomposition `j`, such that `s = j.toSignedMeasure`. This property is known as the Jordan decomposition diff --git a/Mathlib/MeasureTheory/Function/Egorov.lean b/Mathlib/MeasureTheory/Function/Egorov.lean index 932a6fefb1d33..93797f4c9e62b 100644 --- a/Mathlib/MeasureTheory/Function/Egorov.lean +++ b/Mathlib/MeasureTheory/Function/Egorov.lean @@ -42,7 +42,7 @@ This definition is useful for Egorov's theorem. -/ def notConvergentSeq [Preorder ι] (f : ι → α → β) (g : α → β) (n : ℕ) (j : ι) : Set α := ⋃ (k) (_ : j ≤ k), { x | 1 / (n + 1 : ℝ) < dist (f k x) (g x) } -variable {n : ℕ} {i j : ι} {s : Set α} {ε : ℝ} {f : ι → α → β} {g : α → β} +variable {n : ℕ} {j : ι} {s : Set α} {ε : ℝ} {f : ι → α → β} {g : α → β} theorem mem_notConvergentSeq_iff [Preorder ι] {x : α} : x ∈ notConvergentSeq f g n j ↔ ∃ k ≥ j, 1 / (n + 1 : ℝ) < dist (f k x) (g x) := by @@ -171,7 +171,7 @@ theorem tendstoUniformlyOn_diff_iUnionNotConvergentSeq (hε : 0 < ε) end Egorov -variable [SemilatticeSup ι] [Nonempty ι] [Countable ι] {γ : Type*} [TopologicalSpace γ] +variable [SemilatticeSup ι] [Nonempty ι] [Countable ι] {f : ι → α → β} {g : α → β} {s : Set α} /-- **Egorov's theorem**: If `f : ι → α → β` is a sequence of strongly measurable functions that diff --git a/Mathlib/MeasureTheory/Integral/LebesgueNormedSpace.lean b/Mathlib/MeasureTheory/Integral/LebesgueNormedSpace.lean index 874a14fc17cc0..02feec70974ce 100644 --- a/Mathlib/MeasureTheory/Integral/LebesgueNormedSpace.lean +++ b/Mathlib/MeasureTheory/Integral/LebesgueNormedSpace.lean @@ -13,7 +13,7 @@ open MeasureTheory Filter ENNReal Set open NNReal ENNReal -variable {α β γ δ : Type*} {m : MeasurableSpace α} {μ : MeasureTheory.Measure α} +variable {α : Type*} {m : MeasurableSpace α} {μ : MeasureTheory.Measure α} theorem aemeasurable_withDensity_iff {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] [SecondCountableTopology E] [MeasurableSpace E] [BorelSpace E] {f : α → ℝ≥0} diff --git a/Mathlib/MeasureTheory/Measure/Complex.lean b/Mathlib/MeasureTheory/Measure/Complex.lean index 35a76272bc537..611174a941c5f 100644 --- a/Mathlib/MeasureTheory/Measure/Complex.lean +++ b/Mathlib/MeasureTheory/Measure/Complex.lean @@ -34,7 +34,7 @@ noncomputable section open scoped MeasureTheory ENNReal NNReal -variable {α β : Type*} {m : MeasurableSpace α} +variable {α : Type*} {m : MeasurableSpace α} namespace MeasureTheory diff --git a/Mathlib/MeasureTheory/Measure/VectorMeasure.lean b/Mathlib/MeasureTheory/Measure/VectorMeasure.lean index 7a254f9b3708c..d6bc3aa0248e0 100644 --- a/Mathlib/MeasureTheory/Measure/VectorMeasure.lean +++ b/Mathlib/MeasureTheory/Measure/VectorMeasure.lean @@ -698,7 +698,7 @@ instance instPartialOrder : PartialOrder (VectorMeasure α M) where le_trans _ _ _ h₁ h₂ i hi := le_trans (h₁ i hi) (h₂ i hi) le_antisymm _ _ h₁ h₂ := ext fun i hi => le_antisymm (h₁ i hi) (h₂ i hi) -variable {u v w : VectorMeasure α M} +variable {v w : VectorMeasure α M} theorem le_iff : v ≤ w ↔ ∀ i, MeasurableSet i → v i ≤ w i := Iff.rfl diff --git a/Mathlib/Order/UpperLower/Basic.lean b/Mathlib/Order/UpperLower/Basic.lean index 1c2dbcf53ed26..09c1126722f1f 100644 --- a/Mathlib/Order/UpperLower/Basic.lean +++ b/Mathlib/Order/UpperLower/Basic.lean @@ -51,7 +51,7 @@ variable {α β γ : Type*} {ι : Sort*} {κ : ι → Sort*} section LE -variable [LE α] [LE β] {s t : Set α} {a : α} +variable [LE α] {s t : Set α} {a : α} /-- An upper set in an order `α` is a set such that any element greater than one of its members is also a member. Also called up-set, upward-closed set. -/ @@ -984,7 +984,7 @@ end UpperSet namespace LowerSet -variable {f : α ≃o β} {s t : LowerSet α} {a : α} {b : β} +variable {f : α ≃o β} {s t : LowerSet α} {a : α} /-- An order isomorphism of Preorders induces an order isomorphism of their lower sets. -/ def map (f : α ≃o β) : LowerSet α ≃o LowerSet β where @@ -1602,7 +1602,7 @@ variable [Preorder α] [Preorder β] section -variable {s : Set α} {t : Set β} {x : α × β} +variable {s : Set α} {t : Set β} theorem IsUpperSet.prod (hs : IsUpperSet s) (ht : IsUpperSet t) : IsUpperSet (s ×ˢ t) := fun _ _ h ha => ⟨hs h.1 ha.1, ht h.2 ha.2⟩ diff --git a/Mathlib/Probability/Kernel/Disintegration/MeasurableStieltjes.lean b/Mathlib/Probability/Kernel/Disintegration/MeasurableStieltjes.lean index fce06642b56a2..9ecc3f1b58b8d 100644 --- a/Mathlib/Probability/Kernel/Disintegration/MeasurableStieltjes.lean +++ b/Mathlib/Probability/Kernel/Disintegration/MeasurableStieltjes.lean @@ -48,7 +48,7 @@ open scoped NNReal ENNReal MeasureTheory Topology namespace ProbabilityTheory -variable {α β ι : Type*} +variable {α : Type*} section IsMeasurableRatCDF diff --git a/Mathlib/Probability/ProbabilityMassFunction/Basic.lean b/Mathlib/Probability/ProbabilityMassFunction/Basic.lean index 877ecc41e7f9d..a28a889e12f1a 100644 --- a/Mathlib/Probability/ProbabilityMassFunction/Basic.lean +++ b/Mathlib/Probability/ProbabilityMassFunction/Basic.lean @@ -32,7 +32,7 @@ probability mass function, discrete probability measure noncomputable section -variable {α β γ : Type*} +variable {α : Type*} open scoped Classical open NNReal ENNReal MeasureTheory @@ -132,7 +132,7 @@ open MeasureTheory MeasureTheory.OuterMeasure def toOuterMeasure (p : PMF α) : OuterMeasure α := OuterMeasure.sum fun x : α => p x • dirac x -variable (p : PMF α) (s t : Set α) +variable (p : PMF α) (s : Set α) theorem toOuterMeasure_apply : p.toOuterMeasure s = ∑' x, s.indicator p x := tsum_congr fun x => smul_dirac_apply (p x) x s @@ -210,7 +210,7 @@ open MeasureTheory def toMeasure [MeasurableSpace α] (p : PMF α) : Measure α := p.toOuterMeasure.toMeasure ((toOuterMeasure_caratheodory p).symm ▸ le_top) -variable [MeasurableSpace α] (p : PMF α) (s t : Set α) +variable [MeasurableSpace α] (p : PMF α) (s : Set α) theorem toOuterMeasure_apply_le_toMeasure_apply : p.toOuterMeasure s ≤ p.toMeasure s := le_toMeasure_apply p.toOuterMeasure _ s @@ -333,8 +333,7 @@ instance toMeasure.isProbabilityMeasure [MeasurableSpace α] (p : PMF α) : simpa only [MeasurableSet.univ, toMeasure_apply_eq_toOuterMeasure_apply, Set.indicator_univ, toOuterMeasure_apply, ENNReal.coe_eq_one] using tsum_coe p⟩ -variable [Countable α] [MeasurableSpace α] [MeasurableSingletonClass α] (p : PMF α) (μ : Measure α) - [IsProbabilityMeasure μ] +variable [Countable α] [MeasurableSpace α] [MeasurableSingletonClass α] (p : PMF α) @[simp] theorem toMeasure_toPMF : p.toMeasure.toPMF = p := diff --git a/Mathlib/Topology/ContinuousMap/StoneWeierstrass.lean b/Mathlib/Topology/ContinuousMap/StoneWeierstrass.lean index 0bb2fd98c0bd2..8097bb16a5cf8 100644 --- a/Mathlib/Topology/ContinuousMap/StoneWeierstrass.lean +++ b/Mathlib/Topology/ContinuousMap/StoneWeierstrass.lean @@ -453,7 +453,7 @@ end PolynomialFunctions section ContinuousMapZero -variable {X : Type*} [TopologicalSpace X] {𝕜 : Type*} [RCLike 𝕜] +variable {𝕜 : Type*} [RCLike 𝕜] open NonUnitalStarAlgebra Submodule namespace ContinuousMap diff --git a/Mathlib/Topology/MetricSpace/Pseudo/Basic.lean b/Mathlib/Topology/MetricSpace/Pseudo/Basic.lean index 550dee230352d..94125fb3b8a4c 100644 --- a/Mathlib/Topology/MetricSpace/Pseudo/Basic.lean +++ b/Mathlib/Topology/MetricSpace/Pseudo/Basic.lean @@ -18,9 +18,9 @@ Further results about pseudo-metric spaces. open Set Filter TopologicalSpace Bornology open scoped ENNReal NNReal Uniformity Topology -universe u v w +universe u v -variable {α : Type u} {β : Type v} {X ι : Type*} +variable {α : Type u} {β : Type v} {ι : Type*} variable [PseudoMetricSpace α] @@ -59,7 +59,6 @@ theorem dist_le_range_sum_of_dist_le {f : ℕ → α} (n : ℕ) {d : ℕ → ℝ namespace Metric -- instantiate pseudometric space as a topology -variable {x y z : α} {δ ε ε₁ ε₂ : ℝ} {s : Set α} nonrec theorem isUniformInducing_iff [PseudoMetricSpace β] {f : α → β} : IsUniformInducing f ↔ UniformContinuous f ∧ @@ -163,6 +162,8 @@ protected theorem cauchy_iff {f : Filter α} : Cauchy f ↔ NeBot f ∧ ∀ ε > 0, ∃ t ∈ f, ∀ x ∈ t, ∀ y ∈ t, dist x y < ε := uniformity_basis_dist.cauchy_iff +variable {s : Set α} + /-- Given a point `x` in a discrete subset `s` of a pseudometric space, there is an open ball centered at `x` and intersecting `s` only at `x`. -/ theorem exists_ball_inter_eq_singleton_of_mem_discrete [DiscreteTopology s] {x : α} (hx : x ∈ s) : @@ -206,8 +207,6 @@ end Real namespace Metric -variable {x y z : α} {ε ε₁ ε₂ : ℝ} {s : Set α} - -- Porting note: `TopologicalSpace.IsSeparable.separableSpace` moved to `EMetricSpace` /-- The preimage of a separable set by an inducing map is separable. -/ From ede775853cc9f271a9fcf1475dec887f82544e30 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Wed, 16 Oct 2024 18:15:31 +0000 Subject: [PATCH 037/505] =?UTF-8?q?feat(AlgebraicGeometry):=20Classificati?= =?UTF-8?q?on=20of=20`Spec=20R=20=E2=9F=B6=20X`=20with=20`R`=20local.=20(#?= =?UTF-8?q?15240)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- .../PrimeSpectrum/Basic.lean | 4 + Mathlib/AlgebraicGeometry/Scheme.lean | 11 + Mathlib/AlgebraicGeometry/Stalk.lean | 270 +++++++++++++++++- 3 files changed, 276 insertions(+), 9 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean index f3fa7a4fbdf4a..26c8303a68816 100644 --- a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean +++ b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean @@ -647,6 +647,10 @@ theorem closedPoint_mem_iff (U : TopologicalSpace.Opens <| PrimeSpectrum R) : · rintro rfl trivial +lemma closed_point_mem_iff {U : TopologicalSpace.Opens (PrimeSpectrum R)} : + closedPoint R ∈ U ↔ U = ⊤ := + ⟨(eq_top_iff.mpr fun x _ ↦ (specializes_closedPoint x).mem_open U.2 ·), (· ▸ trivial)⟩ + @[simp] theorem PrimeSpectrum.comap_residue (T : Type u) [CommRing T] [LocalRing T] (x : PrimeSpectrum (ResidueField T)) : PrimeSpectrum.comap (residue T) x = closedPoint T := by diff --git a/Mathlib/AlgebraicGeometry/Scheme.lean b/Mathlib/AlgebraicGeometry/Scheme.lean index 3a705c863a6a4..546fea88d15d5 100644 --- a/Mathlib/AlgebraicGeometry/Scheme.lean +++ b/Mathlib/AlgebraicGeometry/Scheme.lean @@ -709,4 +709,15 @@ end Scheme end Stalks +section LocalRing + +open LocalRing + +@[simp] +lemma Spec_closedPoint {R S : CommRingCat} [LocalRing R] [LocalRing S] + {f : R ⟶ S} [IsLocalRingHom f] : (Spec.map f).base (closedPoint S) = closedPoint R := + LocalRing.comap_closedPoint f + +end LocalRing + end AlgebraicGeometry diff --git a/Mathlib/AlgebraicGeometry/Stalk.lean b/Mathlib/AlgebraicGeometry/Stalk.lean index 8865ed5e23d8f..c5d01c4906552 100644 --- a/Mathlib/AlgebraicGeometry/Stalk.lean +++ b/Mathlib/AlgebraicGeometry/Stalk.lean @@ -8,14 +8,25 @@ import Mathlib.AlgebraicGeometry.AffineScheme /-! # Stalks of a Scheme -Given a scheme `X` and a point `x : X`, `AlgebraicGeometry.Scheme.fromSpecStalk X x` is the -canonical scheme morphism from `Spec(O_x)` to `X`. This is helpful for constructing the canonical -map from the spectrum of the residue field of a point to the original scheme. +## Main definitions and results + +- `AlgebraicGeometry.Scheme.fromSpecStalk`: The canonical morphism `Spec 𝒪_{X, x} ⟶ X`. +- `AlgebraicGeometry.Scheme.range_fromSpecStalk`: The range of the map `Spec 𝒪_{X, x} ⟶ X` is + exactly the `y`s that specialize to `x`. +- `AlgebraicGeometry.SpecToEquivOfLocalRing`: + Given a local ring `R` and scheme `X`, morphisms `Spec R ⟶ X` corresponds to pairs + `(x, f)` where `x : X` and `f : 𝒪_{X, x} ⟶ R` is a local ring homomorphism. -/ namespace AlgebraicGeometry -open CategoryTheory Opposite TopologicalSpace +open CategoryTheory Opposite TopologicalSpace LocalRing + +universe u + +variable {X Y : Scheme.{u}} (f : X ⟶ Y) {U V : X.Opens} (hU : IsAffineOpen U) (hV : IsAffineOpen V) + +section fromSpecStalk /-- A morphism from `Spec(O_x)` to `X`, which is defined with the help of an affine open @@ -30,8 +41,7 @@ noncomputable def IsAffineOpen.fromSpecStalk The morphism from `Spec(O_x)` to `X` given by `IsAffineOpen.fromSpec` does not depend on the affine open neighborhood of `x` we choose. -/ -theorem IsAffineOpen.fromSpecStalk_eq {X : Scheme} (x : X) {U V : X.Opens} - (hU : IsAffineOpen U) (hV : IsAffineOpen V) (hxU : x ∈ U) (hxV : x ∈ V) : +theorem IsAffineOpen.fromSpecStalk_eq (x : X) (hxU : x ∈ U) (hxV : x ∈ V) : hU.fromSpecStalk hxU = hV.fromSpecStalk hxV := by obtain ⟨U', h₁, h₂, h₃ : U' ≤ U ⊓ V⟩ := Opens.isBasis_iff_nbhd.mp (isBasis_affine_open X) (show x ∈ U ⊓ V from ⟨hxU, hxV⟩) @@ -51,12 +61,254 @@ theorem IsAffineOpen.fromSpecStalk_eq {X : Scheme} (x : X) {U V : X.Opens} If `x` is a point of `X`, this is the canonical morphism from `Spec(O_x)` to `X`. -/ noncomputable def Scheme.fromSpecStalk (X : Scheme) (x : X) : - Scheme.Spec.obj (op (X.presheaf.stalk x)) ⟶ X := + Spec (X.presheaf.stalk x) ⟶ X := (isAffineOpen_opensRange (X.affineOpenCover.map x)).fromSpecStalk (X.affineOpenCover.covers x) @[simp] -theorem IsAffineOpen.fromSpecStalk_eq_fromSpecStalk - {X : Scheme} {U : X.Opens} (hU : IsAffineOpen U) {x : X} (hxU : x ∈ U) : +theorem IsAffineOpen.fromSpecStalk_eq_fromSpecStalk {x : X} (hxU : x ∈ U) : hU.fromSpecStalk hxU = X.fromSpecStalk x := fromSpecStalk_eq .. +lemma IsAffineOpen.fromSpecStalk_closedPoint {U : Opens X} (hU : IsAffineOpen U) + {x : X} (hxU : x ∈ U) : + (hU.fromSpecStalk hxU).base (closedPoint (X.presheaf.stalk x)) = x := by + rw [IsAffineOpen.fromSpecStalk, Scheme.comp_base_apply] + rw [← hU.primeIdealOf_eq_map_closedPoint ⟨x, hxU⟩, hU.fromSpec_primeIdealOf ⟨x, hxU⟩] + +namespace Scheme + +@[simp] +lemma fromSpecStalk_closedPoint {x : X} : + (X.fromSpecStalk x).base (closedPoint (X.presheaf.stalk x)) = x := + IsAffineOpen.fromSpecStalk_closedPoint _ _ + +lemma fromSpecStalk_app {x : X} (hxU : x ∈ U) : + (X.fromSpecStalk x).app U = + X.presheaf.germ U x hxU ≫ + (ΓSpecIso (X.presheaf.stalk x)).inv ≫ + (Spec (X.presheaf.stalk x)).presheaf.map (homOfLE le_top).op := by + obtain ⟨_, ⟨V : X.Opens, hV, rfl⟩, hxV, hVU⟩ := (isBasis_affine_open X).exists_subset_of_mem_open + hxU U.2 + rw [← hV.fromSpecStalk_eq_fromSpecStalk hxV, IsAffineOpen.fromSpecStalk, Scheme.comp_app, + hV.fromSpec_app_of_le _ hVU, ← X.presheaf.germ_res (homOfLE hVU) x hxV] + simp [Category.assoc, ← ΓSpecIso_inv_naturality_assoc] + +@[reassoc] +lemma Spec_map_stalkSpecializes_fromSpecStalk {x y : X} (h : x ⤳ y) : + Spec.map (X.presheaf.stalkSpecializes h) ≫ X.fromSpecStalk y = X.fromSpecStalk x := by + obtain ⟨_, ⟨U, hU, rfl⟩, hyU, -⟩ := + (isBasis_affine_open X).exists_subset_of_mem_open (Set.mem_univ y) isOpen_univ + have hxU : x ∈ U := h.mem_open U.2 hyU + rw [← hU.fromSpecStalk_eq_fromSpecStalk hyU, ← hU.fromSpecStalk_eq_fromSpecStalk hxU, + IsAffineOpen.fromSpecStalk, IsAffineOpen.fromSpecStalk, ← Category.assoc, ← Spec.map_comp, + TopCat.Presheaf.germ_stalkSpecializes] + +@[reassoc (attr := simp)] +lemma Spec_map_stalkMap_fromSpecStalk {x} : + Spec.map (f.stalkMap x) ≫ Y.fromSpecStalk _ = X.fromSpecStalk x ≫ f := by + obtain ⟨_, ⟨U, hU, rfl⟩, hxU, -⟩ := (isBasis_affine_open Y).exists_subset_of_mem_open + (Set.mem_univ (f.base x)) isOpen_univ + obtain ⟨_, ⟨V, hV, rfl⟩, hxV, hVU⟩ := (isBasis_affine_open X).exists_subset_of_mem_open + hxU (f ⁻¹ᵁ U).2 + rw [← hU.fromSpecStalk_eq_fromSpecStalk hxU, ← hV.fromSpecStalk_eq_fromSpecStalk hxV, + IsAffineOpen.fromSpecStalk, ← Spec.map_comp_assoc, Scheme.stalkMap_germ f _ x hxU, + IsAffineOpen.fromSpecStalk, Spec.map_comp_assoc, ← X.presheaf.germ_res (homOfLE hVU) x hxV, + Spec.map_comp_assoc, Category.assoc, ← Spec.map_comp_assoc (f.app _), + Hom.app_eq_appLE, Hom.appLE_map, IsAffineOpen.Spec_map_appLE_fromSpec] + +lemma Spec_fromSpecStalk (R : CommRingCat) (x) : + (Spec R).fromSpecStalk x = + Spec.map ((ΓSpecIso R).inv ≫ (Spec R).presheaf.germ ⊤ x trivial) := by + rw [← (isAffineOpen_top (Spec R)).fromSpecStalk_eq_fromSpecStalk (x := x) trivial, + IsAffineOpen.fromSpecStalk, IsAffineOpen.fromSpec_top, isoSpec_Spec_inv, + ← Spec.map_comp] + +-- This is not a simp lemma to respect the abstraction boundaries +/-- A variant of `Spec_fromSpecStalk` that breaks abstraction boundaries. -/ +lemma Spec_fromSpecStalk' (R : CommRingCat) (x) : + (Spec R).fromSpecStalk x = Spec.map (StructureSheaf.toStalk R _) := + Spec_fromSpecStalk _ _ + +@[stacks 01J7] +lemma range_fromSpecStalk {x : X} : + Set.range (X.fromSpecStalk x).base = { y | y ⤳ x } := by + ext y + constructor + · rintro ⟨y, rfl⟩ + exact ((LocalRing.specializes_closedPoint y).map (X.fromSpecStalk x).base.2).trans + (specializes_of_eq fromSpecStalk_closedPoint) + · rintro (hy : y ⤳ x) + have := fromSpecStalk_closedPoint (x := y) + rw [← Spec_map_stalkSpecializes_fromSpecStalk hy] at this + exact ⟨_, this⟩ + +end Scheme + +end fromSpecStalk + +variable (R : CommRingCat.{u}) [LocalRing R] + +section stalkClosedPointIso + +/-- For a local ring `(R, 𝔪)`, +this is the isomorphism between the stalk of `Spec R` at `𝔪` and `R`. -/ +noncomputable +def stalkClosedPointIso : + (Spec R).presheaf.stalk (closedPoint R) ≅ R := + StructureSheaf.stalkIso _ _ ≪≫ (IsLocalization.atUnits R + (closedPoint R).asIdeal.primeCompl fun _ ↦ not_not.mp).toRingEquiv.toCommRingCatIso.symm + +lemma stalkClosedPointIso_inv : + (stalkClosedPointIso R).inv = StructureSheaf.toStalk R _ := by + ext x + exact StructureSheaf.localizationToStalk_of _ _ _ + +lemma ΓSpecIso_hom_stalkClosedPointIso_inv : + (Scheme.ΓSpecIso R).hom ≫ (stalkClosedPointIso R).inv = + (Spec R).presheaf.germ ⊤ (closedPoint _) trivial := by + rw [stalkClosedPointIso_inv, ← Iso.eq_inv_comp] + rfl + +@[reassoc (attr := simp)] +lemma germ_stalkClosedPointIso_hom : + (Spec R).presheaf.germ ⊤ (closedPoint _) trivial ≫ (stalkClosedPointIso R).hom = + (Scheme.ΓSpecIso R).hom := by + rw [← ΓSpecIso_hom_stalkClosedPointIso_inv, Category.assoc, Iso.inv_hom_id, Category.comp_id] + +lemma Spec_stalkClosedPointIso : + Spec.map (stalkClosedPointIso R).inv = (Spec R).fromSpecStalk (closedPoint R) := by + rw [stalkClosedPointIso_inv, Scheme.Spec_fromSpecStalk'] + +end stalkClosedPointIso + +section stalkClosedPointTo + +variable {R} (f : Spec R ⟶ X) + + +namespace Scheme + +/-- +Given a local ring `(R, 𝔪)` and a morphism `f : Spec R ⟶ X`, +they induce a (local) ring homomorphism `φ : 𝒪_{X, f 𝔪} ⟶ R`. + +This is inverse to `φ ↦ Spec.map φ ≫ X.fromSpecStalk (f 𝔪)`. See `SpecToEquivOfLocalRing`. +-/ +noncomputable +def stalkClosedPointTo : + X.presheaf.stalk (f.base (closedPoint R)) ⟶ R := + f.stalkMap (closedPoint R) ≫ (stalkClosedPointIso R).hom + +instance isLocalRingHom_stalkClosedPointTo : + IsLocalRingHom (stalkClosedPointTo f) := by + apply (config := { allowSynthFailures := true }) RingHom.isLocalRingHom_comp + · apply isLocalRingHom_of_iso + · apply f.prop + +lemma preimage_eq_top_of_closedPoint_mem + {U : Opens X} (hU : f.base (closedPoint R) ∈ U) : f ⁻¹ᵁ U = ⊤ := + LocalRing.closed_point_mem_iff.mp hU + +lemma stalkClosedPointTo_comp (g : X ⟶ Y) : + stalkClosedPointTo (f ≫ g) = g.stalkMap _ ≫ stalkClosedPointTo f := by + rw [stalkClosedPointTo, Scheme.stalkMap_comp] + exact Category.assoc _ _ _ + +lemma germ_stalkClosedPointTo_Spec {R S : CommRingCat} [LocalRing S] (φ : R ⟶ S): + (Spec R).presheaf.germ ⊤ _ trivial ≫ stalkClosedPointTo (Spec.map φ) = + (ΓSpecIso R).hom ≫ φ := by + rw [stalkClosedPointTo, Scheme.stalkMap_germ_assoc, ← Iso.inv_comp_eq, + ← ΓSpecIso_inv_naturality_assoc] + simp_rw [Opens.map_top] + rw [germ_stalkClosedPointIso_hom, Iso.inv_hom_id, Category.comp_id] + +@[reassoc] +lemma germ_stalkClosedPointTo (U : Opens X) (hU : f.base (closedPoint R) ∈ U) : + X.presheaf.germ U _ hU ≫ stalkClosedPointTo f = f.app U ≫ + ((Spec R).presheaf.mapIso (eqToIso (preimage_eq_top_of_closedPoint_mem f hU).symm).op ≪≫ + ΓSpecIso R).hom := by + rw [stalkClosedPointTo, Scheme.stalkMap_germ_assoc, Iso.trans_hom] + congr 1 + rw [← Iso.eq_comp_inv, Category.assoc, ΓSpecIso_hom_stalkClosedPointIso_inv] + simp only [TopCat.Presheaf.pushforward_obj_obj, Functor.mapIso_hom, Iso.op_hom, eqToIso.hom, + TopCat.Presheaf.germ_res] + +@[reassoc] +lemma germ_stalkClosedPointTo_Spec_fromSpecStalk + {x : X} (f : X.presheaf.stalk x ⟶ R) [IsLocalRingHom f] (U : Opens X) (hU) : + X.presheaf.germ U _ hU ≫ stalkClosedPointTo (Spec.map f ≫ X.fromSpecStalk x) = + X.presheaf.germ U x (by simpa using hU) ≫ f := by + have : (Spec.map f ≫ X.fromSpecStalk x).base (closedPoint R) = x := by + rw [comp_base_apply, Spec_closedPoint, fromSpecStalk_closedPoint] + have : x ∈ U := this ▸ hU + simp only [TopCat.Presheaf.stalkCongr_hom, TopCat.Presheaf.germ_stalkSpecializes_assoc, + germ_stalkClosedPointTo, comp_app, + fromSpecStalk_app (X := X) (x := x) this, Category.assoc, Iso.trans_hom, + Functor.mapIso_hom, Hom.naturality_assoc, ← Functor.map_comp_assoc, + (Spec.map f).app_eq_appLE, Hom.appLE_map_assoc, Hom.map_appLE_assoc] + simp_rw [← Opens.map_top (Spec.map f).base] + rw [← (Spec.map f).app_eq_appLE, ΓSpecIso_naturality, Iso.inv_hom_id_assoc] + +lemma stalkClosedPointTo_fromSpecStalk (x : X) : + stalkClosedPointTo (X.fromSpecStalk x) = + (X.presheaf.stalkCongr (by rw [fromSpecStalk_closedPoint]; rfl)).hom := by + refine TopCat.Presheaf.stalk_hom_ext _ fun U hxU ↦ ?_ + simp only [TopCat.Presheaf.stalkCongr_hom, TopCat.Presheaf.germ_stalkSpecializes, id_eq] + have : X.fromSpecStalk x = Spec.map (𝟙 (X.presheaf.stalk x)) ≫ X.fromSpecStalk x := by simp + convert germ_stalkClosedPointTo_Spec_fromSpecStalk (𝟙 (X.presheaf.stalk x)) U hxU + +@[reassoc] +lemma Spec_stalkClosedPointTo_fromSpecStalk : + Spec.map (stalkClosedPointTo f) ≫ X.fromSpecStalk _ = f := by + obtain ⟨_, ⟨U, hU, rfl⟩, hxU, -⟩ := (isBasis_affine_open X).exists_subset_of_mem_open + (Set.mem_univ (f.base (closedPoint R))) isOpen_univ + have := IsAffineOpen.Spec_map_appLE_fromSpec f hU (isAffineOpen_top _) + (preimage_eq_top_of_closedPoint_mem f hxU).ge + rw [IsAffineOpen.fromSpec_top, Iso.eq_inv_comp, isoSpec_Spec_hom] at this + rw [← hU.fromSpecStalk_eq_fromSpecStalk hxU, IsAffineOpen.fromSpecStalk, ← Spec.map_comp_assoc, + germ_stalkClosedPointTo] + simpa only [Iso.trans_hom, Functor.mapIso_hom, Iso.op_hom, Category.assoc, + Hom.app_eq_appLE, Hom.appLE_map_assoc, Spec.map_comp_assoc] + +end Scheme + +end stalkClosedPointTo + +variable {R} + +omit [LocalRing R] in +/-- useful lemma for applications of `SpecToEquivOfLocalRing` -/ +lemma SpecToEquivOfLocalRing_eq_iff + {f₁ f₂ : Σ x, { f : X.presheaf.stalk x ⟶ R // IsLocalRingHom f }} : + f₁ = f₂ ↔ ∃ h₁ : f₁.1 = f₂.1, f₁.2.1 = + (X.presheaf.stalkCongr (by rw [h₁]; rfl)).hom ≫ f₂.2.1 := by + constructor + · rintro rfl; simp + · obtain ⟨x₁, ⟨f₁, h₁⟩⟩ := f₁ + obtain ⟨x₂, ⟨f₂, h₂⟩⟩ := f₂ + rintro ⟨rfl : x₁ = x₂, e : f₁ = _⟩ + simp [e] + +variable (X R) + +/-- +Given a local ring `R` and scheme `X`, morphisms `Spec R ⟶ X` corresponds to pairs +`(x, f)` where `x : X` and `f : 𝒪_{X, x} ⟶ R` is a local ring homomorphism. +-/ +@[simps] +noncomputable +def SpecToEquivOfLocalRing : + (Spec R ⟶ X) ≃ Σ x, { f : X.presheaf.stalk x ⟶ R // IsLocalRingHom f } where + toFun f := ⟨f.base (closedPoint R), Scheme.stalkClosedPointTo f, inferInstance⟩ + invFun xf := Spec.map xf.2.1 ≫ X.fromSpecStalk xf.1 + left_inv := Scheme.Spec_stalkClosedPointTo_fromSpecStalk + right_inv xf := by + obtain ⟨x, ⟨f, hf⟩⟩ := xf + symm + refine SpecToEquivOfLocalRing_eq_iff.mpr ⟨?_, ?_⟩ + · simp only [Scheme.comp_coeBase, TopCat.coe_comp, Function.comp_apply, Spec_closedPoint, + Scheme.fromSpecStalk_closedPoint] + · refine TopCat.Presheaf.stalk_hom_ext _ fun U hxU ↦ ?_ + simp only [Scheme.germ_stalkClosedPointTo_Spec_fromSpecStalk, + TopCat.Presheaf.stalkCongr_hom, TopCat.Presheaf.germ_stalkSpecializes_assoc] + end AlgebraicGeometry From f4d8a98ef4127d1b1ba22a15494e35d14a03fa0d Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Wed, 16 Oct 2024 21:14:24 +0000 Subject: [PATCH 038/505] chore: copy `Finsupp.smul_single'` to (`Add`)`MonoidAlgebra` to avoid defeq abuse (#17825) This seems to stem from Lean 3 where it was okay to `simp` with `Finsupp` lemmas in its type synonyms `MonoidAlgebra` and `AddMonoidAlgebra`. The discrimination tree in Lean 4 correctly catches the mismatch. So we copy the lemma to the type synonyms and remove some porting notes. --- Mathlib/Algebra/MonoidAlgebra/Defs.lean | 19 ++++++++++----- Mathlib/Algebra/MonoidAlgebra/Grading.lean | 2 +- .../GroupCohomology/Resolution.lean | 23 ++++--------------- 3 files changed, 18 insertions(+), 26 deletions(-) diff --git a/Mathlib/Algebra/MonoidAlgebra/Defs.lean b/Mathlib/Algebra/MonoidAlgebra/Defs.lean index d7cde54dca700..5f5174af5e945 100644 --- a/Mathlib/Algebra/MonoidAlgebra/Defs.lean +++ b/Mathlib/Algebra/MonoidAlgebra/Defs.lean @@ -477,9 +477,13 @@ def of [MulOneClass G] : G →* MonoidAlgebra k G := end +/-- Copy of `Finsupp.smul_single'` that avoids the `MonoidAlgebra = Finsupp` defeq abuse. -/ +@[simp] +theorem smul_single' (c : k) (a : G) (b : k) : c • single a b = single a (c * b) := + Finsupp.smul_single' c a b + theorem smul_of [MulOneClass G] (g : G) (r : k) : r • of k G g = single g r := by - -- porting note (#10745): was `simp`. - rw [of_apply, smul_single', mul_one] + simp theorem of_injective [MulOneClass G] [Nontrivial k] : Function.Injective (of k G) := fun a b h => by @@ -671,8 +675,7 @@ theorem induction_on [Semiring k] [Monoid G] {p : MonoidAlgebra k G → Prop} (f refine Finsupp.induction_linear f ?_ (fun f g hf hg => hadd f g hf hg) fun g r => ?_ · simpa using hsmul 0 (of k G 1) (hM 1) · convert hsmul r (of k G g) (hM g) - -- Porting note: Was `simp only`. - rw [of_apply, smul_single', mul_one] + simp section @@ -1223,6 +1226,11 @@ def singleHom [AddZeroClass G] : k × Multiplicative G →* k[G] where map_one' := rfl map_mul' _a _b := single_mul_single.symm +/-- Copy of `Finsupp.smul_single'` that avoids the `AddMonoidAlgebra = Finsupp` defeq abuse. -/ +@[simp] +theorem smul_single' (c : k) (a : G) (b : k) : c • single a b = single a (c * b) := + Finsupp.smul_single' c a b + theorem mul_single_apply_aux [Add G] (f : k[G]) (r : k) (x y z : G) (H : ∀ a, a + x = z ↔ a = y) : (f * single x r) z = f y * r := @MonoidAlgebra.mul_single_apply_aux k (Multiplicative G) _ _ _ _ _ _ _ H @@ -1267,8 +1275,7 @@ theorem induction_on [AddMonoid G] {p : k[G] → Prop} (f : k[G]) refine Finsupp.induction_linear f ?_ (fun f g hf hg => hadd f g hf hg) fun g r => ?_ · simpa using hsmul 0 (of k G (Multiplicative.ofAdd 0)) (hM 0) · convert hsmul r (of k G (Multiplicative.ofAdd g)) (hM g) - -- Porting note: Was `simp only`. - rw [of_apply, toAdd_ofAdd, smul_single', mul_one] + simp /-- If `f : G → H` is an additive homomorphism between two additive monoids, then `Finsupp.mapDomain f` is a ring homomorphism between their add monoid algebras. -/ diff --git a/Mathlib/Algebra/MonoidAlgebra/Grading.lean b/Mathlib/Algebra/MonoidAlgebra/Grading.lean index e6e0a806f2d7f..86426ba1b8376 100644 --- a/Mathlib/Algebra/MonoidAlgebra/Grading.lean +++ b/Mathlib/Algebra/MonoidAlgebra/Grading.lean @@ -132,7 +132,7 @@ theorem decomposeAux_single (m : M) (r : R) : refine (DirectSum.of_smul R _ _ _).symm.trans ?_ apply DirectSum.of_eq_of_gradedMonoid_eq refine Sigma.subtype_ext rfl ?_ - refine (Finsupp.smul_single' _ _ _).trans ?_ + refine (smul_single' _ _ _).trans ?_ rw [mul_one] rfl diff --git a/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean b/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean index acca90f2bc60f..c99ba8f5afb4e 100644 --- a/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean +++ b/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean @@ -214,35 +214,20 @@ theorem diagonalSucc_inv_single_left (g : G) (f : Gⁿ →₀ k) (r : k) : (diagonalSucc k G n).inv.hom (Finsupp.single g r ⊗ₜ f) = Finsupp.lift (Gⁿ⁺¹ →₀ k) k Gⁿ (fun f => single (g • partialProd f) r) f := by refine f.induction ?_ ?_ -/- Porting note (#11039): broken proof was · simp only [TensorProduct.tmul_zero, map_zero] - · intro a b x ha hb hx + · intro a b x _ _ hx simp only [lift_apply, smul_single', mul_one, TensorProduct.tmul_add, map_add, diagonalSucc_inv_single_single, hx, Finsupp.sum_single_index, mul_comm b, - zero_mul, single_zero] -/ - · rw [TensorProduct.tmul_zero, map_zero] - rw [map_zero] - · intro _ _ _ _ _ hx - rw [TensorProduct.tmul_add, map_add, map_add, hx] - simp_rw [lift_apply, smul_single, smul_eq_mul] - rw [diagonalSucc_inv_single_single, sum_single_index, mul_comm] - rw [zero_mul, single_zero] + zero_mul, single_zero] theorem diagonalSucc_inv_single_right (g : G →₀ k) (f : Gⁿ) (r : k) : (diagonalSucc k G n).inv.hom (g ⊗ₜ Finsupp.single f r) = Finsupp.lift _ k G (fun a => single (a • partialProd f) r) g := by refine g.induction ?_ ?_ -/- Porting note (#11039): broken proof was · simp only [TensorProduct.zero_tmul, map_zero] - · intro a b x ha hb hx + · intro a b x _ _ hx simp only [lift_apply, smul_single', map_add, hx, diagonalSucc_inv_single_single, - TensorProduct.add_tmul, Finsupp.sum_single_index, zero_mul, single_zero] -/ - · rw [TensorProduct.zero_tmul, map_zero, map_zero] - · intro _ _ _ _ _ hx - rw [TensorProduct.add_tmul, map_add, map_add, hx] - simp_rw [lift_apply, smul_single'] - rw [diagonalSucc_inv_single_single, sum_single_index] - rw [zero_mul, single_zero] + TensorProduct.add_tmul, Finsupp.sum_single_index, zero_mul, single_zero] end Rep From 054aa6cbdca97165c441575e81897f4c1012b7c3 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 16 Oct 2024 22:13:35 +0000 Subject: [PATCH 039/505] feat: `affineSpan k (insert 0 s) = span k s` (#17661) From LeanCamCombi --- .../AffineSpace/AffineSubspace.lean | 18 ++++++++++++++++++ 1 file changed, 18 insertions(+) diff --git a/Mathlib/LinearAlgebra/AffineSpace/AffineSubspace.lean b/Mathlib/LinearAlgebra/AffineSpace/AffineSubspace.lean index 70ea2fa947229..4ad30ae06cd1d 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/AffineSubspace.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/AffineSubspace.lean @@ -49,6 +49,7 @@ noncomputable section open Affine open Set +open scoped Pointwise section @@ -510,6 +511,12 @@ theorem direction_affineSpan (s : Set P) : (affineSpan k s).direction = vectorSp theorem mem_affineSpan {p : P} {s : Set P} (hp : p ∈ s) : p ∈ affineSpan k s := mem_spanPoints k p s hp +@[simp] +lemma vectorSpan_add_self (s : Set V) : (vectorSpan k s : Set V) + s = affineSpan k s := by + ext + simp [mem_add, spanPoints] + aesop + end affineSpan namespace AffineSubspace @@ -1284,6 +1291,17 @@ lemma affineSpan_subset_span {s : Set V} : (affineSpan k s : Set V) ⊆ Submodule.span k s := affineSpan_le_toAffineSubspace_span +-- TODO: We want this to be simp, but `affineSpan` gets simped away to `spanPoints`! +-- Let's delete `spanPoints` +lemma affineSpan_insert_zero (s : Set V) : + (affineSpan k (insert 0 s) : Set V) = Submodule.span k s := by + rw [← Submodule.span_insert_zero] + refine affineSpan_subset_span.antisymm ?_ + rw [← vectorSpan_add_self, vectorSpan_def] + refine Subset.trans ?_ <| subset_add_left _ <| mem_insert .. + gcongr + exact subset_sub_left <| mem_insert .. + end AffineSpace' namespace AffineSubspace From 4061861a77da352e35856873ae5a5be4e99b1fa4 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Thu, 17 Oct 2024 02:22:29 +0000 Subject: [PATCH 040/505] chore: bump proofwidgets (#17848) --- lake-manifest.json | 4 ++-- lakefile.lean | 2 +- 2 files changed, 3 insertions(+), 3 deletions(-) diff --git a/lake-manifest.json b/lake-manifest.json index 6e5aa7951c279..d4790a02b8a62 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -35,10 +35,10 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "cd20dae87c48495f0220663014dff11671597fcf", + "rev": "baa65c6339a56bd22b7292aa4511c54b3cc7a6af", "name": "proofwidgets", "manifestFile": "lake-manifest.json", - "inputRev": "v0.0.43-pre", + "inputRev": "v0.0.43", "inherited": false, "configFile": "lakefile.lean"}, {"url": "https://github.com/leanprover/lean4-cli", diff --git a/lakefile.lean b/lakefile.lean index e0fe819c33970..50c310fac84a6 100644 --- a/lakefile.lean +++ b/lakefile.lean @@ -10,7 +10,7 @@ open Lake DSL require "leanprover-community" / "batteries" @ git "main" require "leanprover-community" / "Qq" @ git "master" require "leanprover-community" / "aesop" @ git "master" -require "leanprover-community" / "proofwidgets" @ git "v0.0.43-pre" +require "leanprover-community" / "proofwidgets" @ git "v0.0.43" require "leanprover-community" / "importGraph" @ git "main" require "leanprover-community" / "LeanSearchClient" @ git "main" from git "https://github.com/leanprover-community/LeanSearchClient" @ "main" From 65e1494f16af0b4c31d71f17e1a784e0659c570a Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 17 Oct 2024 02:29:59 +0000 Subject: [PATCH 041/505] chore: make the measurability assumption be the first argument in `lintegral_indicator` (#17830) ... rather than `f`. Change `integral_indicator` and similar lemmas to match. From GibbsMeasure --- .../Constructions/Prod/Basic.lean | 6 +++--- .../MeasureTheory/Decomposition/Lebesgue.lean | 4 ++-- .../ConditionalExpectation/CondexpL2.lean | 2 +- .../Function/LpSeminorm/Basic.lean | 4 ++-- Mathlib/MeasureTheory/Function/LpSpace.lean | 2 +- .../MeasureTheory/Group/FundamentalDomain.lean | 2 +- Mathlib/MeasureTheory/Group/Prod.lean | 10 +++++----- .../MeasureTheory/Integral/IntegrableOn.lean | 2 +- .../Integral/IntegralEqImproper.lean | 2 +- Mathlib/MeasureTheory/Integral/Layercake.lean | 4 ++-- Mathlib/MeasureTheory/Integral/Lebesgue.lean | 18 +++++++++--------- .../MeasureTheory/Integral/SetIntegral.lean | 2 +- Mathlib/MeasureTheory/Measure/GiryMonad.lean | 2 +- .../Measure/HasOuterApproxClosed.lean | 4 ++-- Mathlib/MeasureTheory/Measure/WithDensity.lean | 2 +- Mathlib/Probability/Distributions/Uniform.lean | 2 +- Mathlib/Probability/Integration.lean | 4 ++-- Mathlib/Probability/Kernel/Composition.lean | 2 +- .../Kernel/Disintegration/Basic.lean | 2 +- .../Kernel/Disintegration/Integral.lean | 4 ++-- .../Kernel/Disintegration/Unique.lean | 2 +- .../Probability/Kernel/MeasureCompProd.lean | 2 +- 22 files changed, 42 insertions(+), 42 deletions(-) diff --git a/Mathlib/MeasureTheory/Constructions/Prod/Basic.lean b/Mathlib/MeasureTheory/Constructions/Prod/Basic.lean index a3700ee3a34cf..a41e36a07b43e 100644 --- a/Mathlib/MeasureTheory/Constructions/Prod/Basic.lean +++ b/Mathlib/MeasureTheory/Constructions/Prod/Basic.lean @@ -260,7 +260,7 @@ theorem Measurable.lintegral_prod_right' [SFinite ν] : ?_ ?_ ?_ · intro c s hs simp only [← indicator_comp_right] - suffices Measurable fun x => c * ν (Prod.mk x ⁻¹' s) by simpa [lintegral_indicator _ (m hs)] + suffices Measurable fun x => c * ν (Prod.mk x ⁻¹' s) by simpa [lintegral_indicator (m hs)] exact (measurable_measure_prod_mk_left hs).const_mul _ · rintro f g - hf - h2f h2g simp only [Pi.add_apply] @@ -333,7 +333,7 @@ theorem prod_prod (s : Set α) (t : Set β) : μ.prod ν (s ×ˢ t) = μ s * ν _ = μ S * ν T := by rw [prod_apply hSTm] simp_rw [mk_preimage_prod_right_eq_if, measure_if, - lintegral_indicator _ (measurableSet_toMeasurable _ _), lintegral_const, + lintegral_indicator (measurableSet_toMeasurable _ _), lintegral_const, restrict_apply_univ, mul_comm] _ = μ s * ν t := by rw [measure_toMeasurable, measure_toMeasurable] · -- Formalization is based on https://mathoverflow.net/a/254134/136589 @@ -899,7 +899,7 @@ theorem lintegral_prod_of_measurable : rw [← indicator_comp_right, const_def, const_comp, ← const_def] conv_rhs => enter [2, x] - rw [lintegral_indicator _ (m (x := x) hs), lintegral_const, + rw [lintegral_indicator (m (x := x) hs), lintegral_const, Measure.restrict_apply MeasurableSet.univ, univ_inter] simp [hs, lintegral_const_mul, measurable_measure_prod_mk_left (ν := ν) hs, prod_apply] · rintro f g - hf _ h2f h2g diff --git a/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean b/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean index 291ea3c448fc5..de3c05de65b39 100644 --- a/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean +++ b/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean @@ -900,7 +900,7 @@ theorem haveLebesgueDecomposition_of_finiteMeasure [IsFiniteMeasure μ] [IsFinit (∫⁻ a in A ∩ E, ε + ξ a ∂ν) + ∫⁻ a in A \ E, ξ a ∂ν := by simp only [lintegral_add_left measurable_const, lintegral_add_left hξm, setLIntegral_const, add_assoc, lintegral_inter_add_diff _ _ hE₁, Pi.add_apply, - lintegral_indicator _ hE₁, restrict_apply hE₁] + lintegral_indicator hE₁, restrict_apply hE₁] rw [inter_comm, add_comm] rw [this, ← measure_inter_add_diff A hE₁] exact add_le_add (hε₂ A hA) (hξle (A \ E) (hA.diff hE₁)) @@ -908,7 +908,7 @@ theorem haveLebesgueDecomposition_of_finiteMeasure [IsFiniteMeasure μ] [IsFinit le_sSup ⟨ξ + E.indicator fun _ ↦ (ε : ℝ≥0∞), hξε, rfl⟩ -- but this contradicts the maximality of `∫⁻ x, ξ x ∂ν` refine not_lt.2 this ?_ - rw [hξ₁, lintegral_add_left hξm, lintegral_indicator _ hE₁, setLIntegral_const] + rw [hξ₁, lintegral_add_left hξm, lintegral_indicator hE₁, setLIntegral_const] refine ENNReal.lt_add_right ?_ (ENNReal.mul_pos_iff.2 ⟨ENNReal.coe_pos.2 hε₁, hE₂⟩).ne' have := measure_ne_top (ν.withDensity ξ) univ rwa [withDensity_apply _ MeasurableSet.univ, Measure.restrict_univ] at this diff --git a/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL2.lean b/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL2.lean index 31ce53180ae05..54f8441a10130 100644 --- a/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL2.lean +++ b/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL2.lean @@ -203,7 +203,7 @@ theorem lintegral_nnnorm_condexpL2_indicator_le_real (hs : MeasurableSet s) (hμ classical simp_rw [Set.indicator_apply] split_ifs <;> simp - rw [h_eq, lintegral_indicator _ hs, lintegral_const, Measure.restrict_restrict hs] + rw [h_eq, lintegral_indicator hs, lintegral_const, Measure.restrict_restrict hs] simp only [one_mul, Set.univ_inter, MeasurableSet.univ, Measure.restrict_apply] end Real diff --git a/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean b/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean index 883a202727ea9..5bda6afb8f9e1 100644 --- a/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean +++ b/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean @@ -335,7 +335,7 @@ theorem eLpNorm_indicator_eq_restrict {f : α → E} {s : Set α} (hs : Measurab ← ENNReal.essSup_indicator_eq_essSup_restrict hs, ENNReal.coe_indicator, nnnorm_indicator_eq_indicator_nnnorm] · rcases eq_or_ne p 0 with rfl | hp₀; · simp - simp only [eLpNorm_eq_lintegral_rpow_nnnorm hp₀ hp, ← lintegral_indicator _ hs, + simp only [eLpNorm_eq_lintegral_rpow_nnnorm hp₀ hp, ← lintegral_indicator hs, ENNReal.coe_indicator, nnnorm_indicator_eq_indicator_nnnorm] congr with x by_cases hx : x ∈ s <;> simp [ENNReal.toReal_pos, *] @@ -1273,7 +1273,7 @@ theorem le_eLpNorm_of_bddBelow (hp : p ≠ 0) (hp' : p ≠ ∞) {f : α → F} ( one_div, ENNReal.le_rpow_inv_iff (ENNReal.toReal_pos hp hp'), ENNReal.mul_rpow_of_nonneg _ _ ENNReal.toReal_nonneg, ← ENNReal.rpow_mul, inv_mul_cancel₀ (ENNReal.toReal_pos hp hp').ne.symm, ENNReal.rpow_one, ← setLIntegral_const, - ← lintegral_indicator _ hs] + ← lintegral_indicator hs] refine lintegral_mono_ae ?_ filter_upwards [hf] with x hx by_cases hxs : x ∈ s diff --git a/Mathlib/MeasureTheory/Function/LpSpace.lean b/Mathlib/MeasureTheory/Function/LpSpace.lean index 12414c825abe8..a20a1c45824a7 100644 --- a/Mathlib/MeasureTheory/Function/LpSpace.lean +++ b/Mathlib/MeasureTheory/Function/LpSpace.lean @@ -616,7 +616,7 @@ theorem eLpNorm_indicator_eq_eLpNorm_restrict {f : α → F} (hs : MeasurableSet simp_rw [eLpNorm_eq_lintegral_rpow_nnnorm hp_zero hp_top] suffices (∫⁻ x, (‖s.indicator f x‖₊ : ℝ≥0∞) ^ p.toReal ∂μ) = ∫⁻ x in s, (‖f x‖₊ : ℝ≥0∞) ^ p.toReal ∂μ by rw [this] - rw [← lintegral_indicator _ hs] + rw [← lintegral_indicator hs] congr simp_rw [nnnorm_indicator_eq_indicator_nnnorm, ENNReal.coe_indicator] have h_zero : (fun x => x ^ p.toReal) (0 : ℝ≥0∞) = 0 := by diff --git a/Mathlib/MeasureTheory/Group/FundamentalDomain.lean b/Mathlib/MeasureTheory/Group/FundamentalDomain.lean index b7e676a68b71b..993b79332bb45 100644 --- a/Mathlib/MeasureTheory/Group/FundamentalDomain.lean +++ b/Mathlib/MeasureTheory/Group/FundamentalDomain.lean @@ -317,7 +317,7 @@ theorem measure_set_eq (hs : IsFundamentalDomain G s μ) (ht : IsFundamentalDoma refine hs.setLIntegral_eq ht (Set.indicator A fun _ => 1) fun g x ↦ ?_ convert (Set.indicator_comp_right (g • · : α → α) (g := fun _ ↦ (1 : ℝ≥0∞))).symm rw [hA g] - simpa [Measure.restrict_apply hA₀, lintegral_indicator _ hA₀] using this + simpa [Measure.restrict_apply hA₀, lintegral_indicator hA₀] using this /-- If `s` and `t` are two fundamental domains of the same action, then their measures are equal. -/ @[to_additive "If `s` and `t` are two fundamental domains of the same action, then their measures diff --git a/Mathlib/MeasureTheory/Group/Prod.lean b/Mathlib/MeasureTheory/Group/Prod.lean index 38b8691d5bc0d..29e7a7bd7f934 100644 --- a/Mathlib/MeasureTheory/Group/Prod.lean +++ b/Mathlib/MeasureTheory/Group/Prod.lean @@ -221,7 +221,7 @@ theorem absolutelyContinuous_map_div_left (g : G) : μ ≪ map (fun h => g / h) @[to_additive "This is the computation performed in the proof of [Halmos, §60 Th. A]."] theorem measure_mul_lintegral_eq [IsMulLeftInvariant ν] (sm : MeasurableSet s) (f : G → ℝ≥0∞) (hf : Measurable f) : (μ s * ∫⁻ y, f y ∂ν) = ∫⁻ x, ν ((fun z => z * x) ⁻¹' s) * f x⁻¹ ∂μ := by - rw [← setLIntegral_one, ← lintegral_indicator _ sm, + rw [← setLIntegral_one, ← lintegral_indicator sm, ← lintegral_lintegral_mul (measurable_const.indicator sm).aemeasurable hf.aemeasurable, ← lintegral_lintegral_mul_inv μ ν] swap @@ -233,7 +233,7 @@ theorem measure_mul_lintegral_eq [IsMulLeftInvariant ν] (sm : MeasurableSet s) have : ∀ x y, s.indicator (fun _ : G => (1 : ℝ≥0∞)) (y * x) = ((fun z => z * x) ⁻¹' s).indicator (fun b : G => 1) y := by intro x y; symm; convert indicator_comp_right (M := ℝ≥0∞) fun y => y * x using 2; ext1; rfl - simp_rw [this, lintegral_mul_const _ (ms _), lintegral_indicator _ (measurable_mul_const _ sm), + simp_rw [this, lintegral_mul_const _ (ms _), lintegral_indicator (measurable_mul_const _ sm), setLIntegral_one] /-- Any two nonzero left-invariant measures are absolutely continuous w.r.t. each other. -/ @@ -262,11 +262,11 @@ theorem ae_measure_preimage_mul_right_lt_top (hμs : μ' s ≠ ∞) : simp only [ν'.inv_apply] at h3A apply ae_lt_top (measurable_measure_mul_right ν' sm) have h1 := measure_mul_lintegral_eq μ' ν' sm (A⁻¹.indicator 1) (measurable_one.indicator hA.inv) - rw [lintegral_indicator _ hA.inv] at h1 + rw [lintegral_indicator hA.inv] at h1 simp_rw [Pi.one_apply, setLIntegral_one, ← image_inv, indicator_image inv_injective, image_inv, ← indicator_mul_right _ fun x => ν' ((fun y => y * x) ⁻¹' s), Function.comp, Pi.one_apply, mul_one] at h1 - rw [← lintegral_indicator _ hA, ← h1] + rw [← lintegral_indicator hA, ← h1] exact ENNReal.mul_ne_top hμs h3A.ne @[to_additive] @@ -320,7 +320,7 @@ theorem measure_mul_measure_eq (s t : Set G) (h2s : ν' s ≠ 0) (h3s : ν' s (measurable_const.indicator ht) have h2 := measure_lintegral_div_measure μ' ν' hs h2s h3s (t.indicator fun _ => 1) (measurable_const.indicator ht) - rw [lintegral_indicator _ ht, setLIntegral_one] at h1 h2 + rw [lintegral_indicator ht, setLIntegral_one] at h1 h2 rw [← h1, mul_left_comm, h2] /-- Left invariant Borel measures on a measurable group are unique (up to a scalar). -/ diff --git a/Mathlib/MeasureTheory/Integral/IntegrableOn.lean b/Mathlib/MeasureTheory/Integral/IntegrableOn.lean index ae1b813c974c1..68ca4ecc4d917 100644 --- a/Mathlib/MeasureTheory/Integral/IntegrableOn.lean +++ b/Mathlib/MeasureTheory/Integral/IntegrableOn.lean @@ -227,7 +227,7 @@ theorem MeasurePreserving.integrableOn_image [MeasurableSpace β] {e : α → β theorem integrable_indicator_iff (hs : MeasurableSet s) : Integrable (indicator s f) μ ↔ IntegrableOn f s μ := by simp [IntegrableOn, Integrable, HasFiniteIntegral, nnnorm_indicator_eq_indicator_nnnorm, - ENNReal.coe_indicator, lintegral_indicator _ hs, aestronglyMeasurable_indicator_iff hs] + ENNReal.coe_indicator, lintegral_indicator hs, aestronglyMeasurable_indicator_iff hs] theorem IntegrableOn.integrable_indicator (h : IntegrableOn f s μ) (hs : MeasurableSet s) : Integrable (indicator s f) μ := diff --git a/Mathlib/MeasureTheory/Integral/IntegralEqImproper.lean b/Mathlib/MeasureTheory/Integral/IntegralEqImproper.lean index d2bec25790e57..bcdccb8f63e40 100644 --- a/Mathlib/MeasureTheory/Integral/IntegralEqImproper.lean +++ b/Mathlib/MeasureTheory/Integral/IntegralEqImproper.lean @@ -361,7 +361,7 @@ private theorem lintegral_tendsto_of_monotone_of_nat {φ : ℕ → Set α} (hφ indicator_le_indicator_of_subset (hmono hij) (fun x => zero_le <| f x) x have key₃ : ∀ᵐ x : α ∂μ, Tendsto (fun n => F n x) atTop (𝓝 (f x)) := hφ.ae_tendsto_indicator f (lintegral_tendsto_of_tendsto_of_monotone key₁ key₂ key₃).congr fun n => - lintegral_indicator f (hφ.measurableSet n) + lintegral_indicator (hφ.measurableSet n) _ theorem AECover.lintegral_tendsto_of_nat {φ : ℕ → Set α} (hφ : AECover μ atTop φ) {f : α → ℝ≥0∞} (hfm : AEMeasurable f μ) : Tendsto (∫⁻ x in φ ·, f x ∂μ) atTop (𝓝 <| ∫⁻ x, f x ∂μ) := by diff --git a/Mathlib/MeasureTheory/Integral/Layercake.lean b/Mathlib/MeasureTheory/Integral/Layercake.lean index ccdbcf47487ab..17be63bb0326b 100644 --- a/Mathlib/MeasureTheory/Integral/Layercake.lean +++ b/Mathlib/MeasureTheory/Integral/Layercake.lean @@ -122,9 +122,9 @@ theorem lintegral_comp_eq_lintegral_meas_le_mul_of_measurable_of_sigmaFinite congr exact intervalIntegral.integral_of_le (f_nn ω) rw [lintegral_congr integrand_eq] - simp_rw [← lintegral_indicator (fun t => ENNReal.ofReal (g t)) measurableSet_Ioc] + simp_rw [← lintegral_indicator measurableSet_Ioc] -- Porting note: was part of `simp_rw` on the previous line, but didn't trigger. - rw [← lintegral_indicator _ measurableSet_Ioi, lintegral_lintegral_swap] + rw [← lintegral_indicator measurableSet_Ioi, lintegral_lintegral_swap] · apply congr_arg funext s have aux₁ : diff --git a/Mathlib/MeasureTheory/Integral/Lebesgue.lean b/Mathlib/MeasureTheory/Integral/Lebesgue.lean index 5be3b2d1f299b..511c11d6a00dd 100644 --- a/Mathlib/MeasureTheory/Integral/Lebesgue.lean +++ b/Mathlib/MeasureTheory/Integral/Lebesgue.lean @@ -757,7 +757,7 @@ theorem lintegral_indicator_le (f : α → ℝ≥0∞) (s : Set α) : simpa [H] using hg x @[simp] -theorem lintegral_indicator (f : α → ℝ≥0∞) {s : Set α} (hs : MeasurableSet s) : +theorem lintegral_indicator {s : Set α} (hs : MeasurableSet s) (f : α → ℝ≥0∞) : ∫⁻ a, s.indicator f a ∂μ = ∫⁻ a in s, f a ∂μ := by apply le_antisymm (lintegral_indicator_le f s) simp only [lintegral, ← restrict_lintegral_eq_lintegral_restrict _ hs, iSup_subtype'] @@ -765,20 +765,20 @@ theorem lintegral_indicator (f : α → ℝ≥0∞) {s : Set α} (hs : Measurabl refine ⟨⟨φ.restrict s, fun x => ?_⟩, le_rfl⟩ simp [hφ x, hs, indicator_le_indicator] -lemma setLIntegral_indicator (f : α → ℝ≥0∞) {s t : Set α} (hs : MeasurableSet s) : +lemma setLIntegral_indicator {s t : Set α} (hs : MeasurableSet s) (f : α → ℝ≥0∞) : ∫⁻ a in t, s.indicator f a ∂μ = ∫⁻ a in s ∩ t, f a ∂μ := by - rw [lintegral_indicator _ hs, Measure.restrict_restrict hs] + rw [lintegral_indicator hs, Measure.restrict_restrict hs] -theorem lintegral_indicator₀ (f : α → ℝ≥0∞) {s : Set α} (hs : NullMeasurableSet s μ) : +theorem lintegral_indicator₀ {s : Set α} (hs : NullMeasurableSet s μ) (f : α → ℝ≥0∞) : ∫⁻ a, s.indicator f a ∂μ = ∫⁻ a in s, f a ∂μ := by rw [← lintegral_congr_ae (indicator_ae_eq_of_ae_eq_set hs.toMeasurable_ae_eq), - lintegral_indicator _ (measurableSet_toMeasurable _ _), + lintegral_indicator (measurableSet_toMeasurable _ _), Measure.restrict_congr_set hs.toMeasurable_ae_eq] lemma setLIntegral_indicator₀ (f : α → ℝ≥0∞) {s t : Set α} (hs : NullMeasurableSet s (μ.restrict t)) : ∫⁻ a in t, s.indicator f a ∂μ = ∫⁻ a in s ∩ t, f a ∂μ := by - rw [lintegral_indicator₀ _ hs, Measure.restrict_restrict₀ hs] + rw [lintegral_indicator₀ hs, Measure.restrict_restrict₀ hs] theorem lintegral_indicator_const_le (s : Set α) (c : ℝ≥0∞) : ∫⁻ a, s.indicator (fun _ => c) a ∂μ ≤ c * μ s := @@ -786,7 +786,7 @@ theorem lintegral_indicator_const_le (s : Set α) (c : ℝ≥0∞) : theorem lintegral_indicator_const₀ {s : Set α} (hs : NullMeasurableSet s μ) (c : ℝ≥0∞) : ∫⁻ a, s.indicator (fun _ => c) a ∂μ = c * μ s := by - rw [lintegral_indicator₀ _ hs, setLIntegral_const] + rw [lintegral_indicator₀ hs, setLIntegral_const] theorem lintegral_indicator_const {s : Set α} (hs : MeasurableSet s) (c : ℝ≥0∞) : ∫⁻ a, s.indicator (fun _ => c) a ∂μ = c * μ s := @@ -872,7 +872,7 @@ lemma lintegral_le_meas {s : Set α} {f : α → ℝ≥0∞} (hf : ∀ a, f a lemma setLIntegral_le_meas {s t : Set α} (hs : MeasurableSet s) {f : α → ℝ≥0∞} (hf : ∀ a ∈ s, a ∈ t → f a ≤ 1) (hf' : ∀ a ∈ s, a ∉ t → f a = 0) : ∫⁻ a in s, f a ∂μ ≤ μ t := by - rw [← lintegral_indicator _ hs] + rw [← lintegral_indicator hs] refine lintegral_le_meas (fun a ↦ ?_) (by aesop) by_cases has : a ∈ s <;> [by_cases hat : a ∈ t; skip] <;> simp [*] @@ -1909,7 +1909,7 @@ theorem lintegral_trim {μ : Measure α} (hm : m ≤ m0) {f : α → ℝ≥0∞} refine @Measurable.ennreal_induction α m (fun f => ∫⁻ a, f a ∂μ.trim hm = ∫⁻ a, f a ∂μ) ?_ ?_ ?_ f hf · intro c s hs - rw [lintegral_indicator _ hs, lintegral_indicator _ (hm s hs), setLIntegral_const, + rw [lintegral_indicator hs, lintegral_indicator (hm s hs), setLIntegral_const, setLIntegral_const] suffices h_trim_s : μ.trim hm s = μ s by rw [h_trim_s] exact trim_measurableSet_eq hm hs diff --git a/Mathlib/MeasureTheory/Integral/SetIntegral.lean b/Mathlib/MeasureTheory/Integral/SetIntegral.lean index 0174fcb3cd94f..76b9b535d4c5c 100644 --- a/Mathlib/MeasureTheory/Integral/SetIntegral.lean +++ b/Mathlib/MeasureTheory/Integral/SetIntegral.lean @@ -1475,7 +1475,7 @@ variable [MeasurableSpace X] [PseudoEMetricSpace X] theorem measure_le_lintegral_thickenedIndicatorAux (μ : Measure X) {E : Set X} (E_mble : MeasurableSet E) (δ : ℝ) : μ E ≤ ∫⁻ x, (thickenedIndicatorAux δ E x : ℝ≥0∞) ∂μ := by convert_to lintegral μ (E.indicator fun _ => (1 : ℝ≥0∞)) ≤ lintegral μ (thickenedIndicatorAux δ E) - · rw [lintegral_indicator _ E_mble] + · rw [lintegral_indicator E_mble] simp only [lintegral_one, Measure.restrict_apply, MeasurableSet.univ, univ_inter] · apply lintegral_mono apply indicator_le_thickenedIndicatorAux diff --git a/Mathlib/MeasureTheory/Measure/GiryMonad.lean b/Mathlib/MeasureTheory/Measure/GiryMonad.lean index 502d9433a781a..5f92026d9e1e7 100644 --- a/Mathlib/MeasureTheory/Measure/GiryMonad.lean +++ b/Mathlib/MeasureTheory/Measure/GiryMonad.lean @@ -182,7 +182,7 @@ theorem dirac_bind {f : α → Measure β} (hf : Measurable f) (a : α) : bind ( @[simp] theorem bind_dirac {m : Measure α} : bind m dirac = m := by ext1 s hs - simp only [bind_apply hs measurable_dirac, dirac_apply' _ hs, lintegral_indicator 1 hs, + simp only [bind_apply hs measurable_dirac, dirac_apply' _ hs, lintegral_indicator hs, Pi.one_apply, lintegral_one, restrict_apply, MeasurableSet.univ, univ_inter] @[simp] diff --git a/Mathlib/MeasureTheory/Measure/HasOuterApproxClosed.lean b/Mathlib/MeasureTheory/Measure/HasOuterApproxClosed.lean index d824d0aecc8ad..6fe5d37bd096f 100644 --- a/Mathlib/MeasureTheory/Measure/HasOuterApproxClosed.lean +++ b/Mathlib/MeasureTheory/Measure/HasOuterApproxClosed.lean @@ -80,7 +80,7 @@ theorem measure_of_cont_bdd_of_tendsto_filter_indicator {ι : Type*} {L : Filter convert tendsto_lintegral_nn_filter_of_le_const μ fs_bdd fs_lim have aux : ∀ ω, indicator E (fun _ ↦ (1 : ℝ≥0∞)) ω = ↑(indicator E (fun _ ↦ (1 : ℝ≥0)) ω) := fun ω ↦ by simp only [ENNReal.coe_indicator, ENNReal.coe_one] - simp_rw [← aux, lintegral_indicator _ E_mble] + simp_rw [← aux, lintegral_indicator E_mble] simp only [lintegral_one, Measure.restrict_apply, MeasurableSet.univ, univ_inter] /-- If a sequence of bounded continuous functions tends to the indicator of a measurable set and @@ -162,7 +162,7 @@ approximating sequence to the indicator of the set. -/ theorem measure_le_lintegral [MeasurableSpace X] [OpensMeasurableSpace X] (μ : Measure X) (n : ℕ) : μ F ≤ ∫⁻ x, (hF.apprSeq n x : ℝ≥0∞) ∂μ := by convert_to ∫⁻ x, (F.indicator (fun _ ↦ (1 : ℝ≥0∞))) x ∂μ ≤ ∫⁻ x, hF.apprSeq n x ∂μ - · rw [lintegral_indicator _ hF.measurableSet] + · rw [lintegral_indicator hF.measurableSet] simp only [lintegral_one, MeasurableSet.univ, Measure.restrict_apply, univ_inter] · apply lintegral_mono intro x diff --git a/Mathlib/MeasureTheory/Measure/WithDensity.lean b/Mathlib/MeasureTheory/Measure/WithDensity.lean index f7b67208ce306..d343307b3e3fc 100644 --- a/Mathlib/MeasureTheory/Measure/WithDensity.lean +++ b/Mathlib/MeasureTheory/Measure/WithDensity.lean @@ -170,7 +170,7 @@ theorem withDensity_tsum {ι : Type*} [Countable ι] {f : ι → α → ℝ≥0 theorem withDensity_indicator {s : Set α} (hs : MeasurableSet s) (f : α → ℝ≥0∞) : μ.withDensity (s.indicator f) = (μ.restrict s).withDensity f := by ext1 t ht - rw [withDensity_apply _ ht, lintegral_indicator _ hs, restrict_comm hs, ← + rw [withDensity_apply _ ht, lintegral_indicator hs, restrict_comm hs, ← withDensity_apply _ ht] theorem withDensity_indicator_one {s : Set α} (hs : MeasurableSet s) : diff --git a/Mathlib/Probability/Distributions/Uniform.lean b/Mathlib/Probability/Distributions/Uniform.lean index fca2b2cafde42..43f54dab8c27d 100644 --- a/Mathlib/Probability/Distributions/Uniform.lean +++ b/Mathlib/Probability/Distributions/Uniform.lean @@ -155,7 +155,7 @@ theorem mul_pdf_integrable (hcs : IsCompact s) (huX : IsUniform X s ℙ) : set ind := (volume s)⁻¹ • (1 : ℝ → ℝ≥0∞) have : ∀ x, ↑‖x‖₊ * s.indicator ind x = s.indicator (fun x => ‖x‖₊ * ind x) x := fun x => (s.indicator_mul_right (fun x => ↑‖x‖₊) ind).symm - simp only [ind, this, lintegral_indicator _ hcs.measurableSet, mul_one, Algebra.id.smul_eq_mul, + simp only [ind, this, lintegral_indicator hcs.measurableSet, mul_one, Algebra.id.smul_eq_mul, Pi.one_apply, Pi.smul_apply] rw [lintegral_mul_const _ measurable_nnnorm.coe_nnreal_ennreal] exact ENNReal.mul_ne_top (setLIntegral_lt_top_of_isCompact hnt.2 hcs continuous_nnnorm).ne diff --git a/Mathlib/Probability/Integration.lean b/Mathlib/Probability/Integration.lean index ad5bc0600a169..3d42e69819c63 100644 --- a/Mathlib/Probability/Integration.lean +++ b/Mathlib/Probability/Integration.lean @@ -51,8 +51,8 @@ theorem lintegral_mul_indicator_eq_lintegral_mul_lintegral_indicator {Mf mΩ : M apply @Measurable.ennreal_induction _ Mf · intro c' s' h_meas_s' simp_rw [← inter_indicator_mul] - rw [lintegral_indicator _ (MeasurableSet.inter (hMf _ h_meas_s') h_meas_T), - lintegral_indicator _ (hMf _ h_meas_s'), lintegral_indicator _ h_meas_T] + rw [lintegral_indicator (MeasurableSet.inter (hMf _ h_meas_s') h_meas_T), + lintegral_indicator (hMf _ h_meas_s'), lintegral_indicator h_meas_T] simp only [measurable_const, lintegral_const, univ_inter, lintegral_const_mul, MeasurableSet.univ, Measure.restrict_apply] rw [IndepSets_iff] at h_ind diff --git a/Mathlib/Probability/Kernel/Composition.lean b/Mathlib/Probability/Kernel/Composition.lean index 7fbfe68d509e1..fd42f53736b76 100644 --- a/Mathlib/Probability/Kernel/Composition.lean +++ b/Mathlib/Probability/Kernel/Composition.lean @@ -323,7 +323,7 @@ theorem compProd_restrict {s : Set β} {t : Set γ} (hs : MeasurableSet s) (ht : · simp only [h, true_and, Set.inter_def, Set.mem_setOf] · simp only [h, false_and, and_false, Set.setOf_false, measure_empty] simp_rw [this] - rw [lintegral_indicator _ hs] + rw [lintegral_indicator hs] theorem compProd_restrict_left {s : Set β} (hs : MeasurableSet s) : Kernel.restrict κ hs ⊗ₖ η = Kernel.restrict (κ ⊗ₖ η) (hs.prod MeasurableSet.univ) := by diff --git a/Mathlib/Probability/Kernel/Disintegration/Basic.lean b/Mathlib/Probability/Kernel/Disintegration/Basic.lean index 0b385aef78547..85c938f894937 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Basic.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Basic.lean @@ -79,7 +79,7 @@ private lemma IsCondKernel.apply_of_ne_zero_of_measurableSet [MeasurableSingleto have : Prod.mk a ⁻¹' (Prod.mk x '' s) = ∅ := by ext y; simp [Ne.symm hax] simp only [this, measure_empty] simp_rw [this] - rw [MeasureTheory.lintegral_indicator _ (measurableSet_singleton x)] + rw [MeasureTheory.lintegral_indicator (measurableSet_singleton x)] simp only [Measure.restrict_singleton, lintegral_smul_measure, lintegral_dirac] rw [← mul_assoc, ENNReal.inv_mul_cancel hx (measure_ne_top _ _), one_mul] diff --git a/Mathlib/Probability/Kernel/Disintegration/Integral.lean b/Mathlib/Probability/Kernel/Disintegration/Integral.lean index 026257adcaf92..6f9b34c32006f 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Integral.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Integral.lean @@ -53,7 +53,7 @@ lemma setLIntegral_condKernel_eq_measure_prod (a : α) {s : Set β} (hs : Measur intro b by_cases hb : b ∈ s <;> simp [hb] simp_rw [this] - rw [lintegral_indicator _ hs] + rw [lintegral_indicator hs] @[deprecated (since := "2024-06-29")] alias set_lintegral_condKernel_eq_measure_prod := setLIntegral_condKernel_eq_measure_prod @@ -172,7 +172,7 @@ lemma setLIntegral_condKernel_eq_measure_prod {s : Set β} (hs : MeasurableSet s intro b by_cases hb : b ∈ s <;> simp [hb] simp_rw [this] - rw [lintegral_indicator _ hs] + rw [lintegral_indicator hs] @[deprecated (since := "2024-06-29")] alias set_lintegral_condKernel_eq_measure_prod := setLIntegral_condKernel_eq_measure_prod diff --git a/Mathlib/Probability/Kernel/Disintegration/Unique.lean b/Mathlib/Probability/Kernel/Disintegration/Unique.lean index e1b79719abfe1..3a2d73da4ca97 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Unique.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Unique.lean @@ -50,7 +50,7 @@ theorem eq_condKernel_of_measure_eq_compProd' (κ : Kernel α Ω) [IsSFiniteKern refine ae_eq_of_forall_setLIntegral_eq_of_sigmaFinite (Kernel.measurable_coe κ hs) (Kernel.measurable_coe ρ.condKernel hs) (fun t ht _ ↦ ?_) conv_rhs => rw [Measure.setLIntegral_condKernel_eq_measure_prod ht hs, hκ] - simp only [Measure.compProd_apply (ht.prod hs), Set.mem_prod, ← lintegral_indicator _ ht] + simp only [Measure.compProd_apply (ht.prod hs), Set.mem_prod, ← lintegral_indicator ht] congr with x by_cases hx : x ∈ t all_goals simp [hx] diff --git a/Mathlib/Probability/Kernel/MeasureCompProd.lean b/Mathlib/Probability/Kernel/MeasureCompProd.lean index b2bf06645a1a6..9e4b40f0b1773 100644 --- a/Mathlib/Probability/Kernel/MeasureCompProd.lean +++ b/Mathlib/Probability/Kernel/MeasureCompProd.lean @@ -61,7 +61,7 @@ lemma compProd_apply [SFinite μ] [IsSFiniteKernel κ] {s : Set (α × β)} (hs lemma compProd_apply_prod [SFinite μ] [IsSFiniteKernel κ] {s : Set α} {t : Set β} (hs : MeasurableSet s) (ht : MeasurableSet t) : (μ ⊗ₘ κ) (s ×ˢ t) = ∫⁻ a in s, κ a t ∂μ := by - rw [compProd_apply (hs.prod ht), ← lintegral_indicator _ hs] + rw [compProd_apply (hs.prod ht), ← lintegral_indicator hs] congr with a classical rw [Set.indicator_apply] From b33faadf563e7b3d6b9578d89a7f04292fc0f86d Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Thu, 17 Oct 2024 02:58:15 +0000 Subject: [PATCH 042/505] chore(Topology/Group): move `QuotientGroup` to a new file (#17473) --- Mathlib.lean | 1 + .../Constructions/Polish/Basic.lean | 5 +- .../MeasureTheory/Measure/Haar/Quotient.lean | 2 +- Mathlib/Topology/Algebra/Group/Basic.lean | 141 +----------------- Mathlib/Topology/Algebra/Group/Quotient.lean | 133 +++++++++++++++++ Mathlib/Topology/Algebra/Module/Basic.lean | 2 +- Mathlib/Topology/Algebra/OpenSubgroup.lean | 1 + .../Topology/Algebra/ProperAction/Basic.lean | 1 + Mathlib/Topology/Algebra/Ring/Ideal.lean | 1 + Mathlib/Topology/Algebra/UniformGroup.lean | 2 +- .../Topology/Compactness/LocallyCompact.lean | 14 ++ 11 files changed, 158 insertions(+), 145 deletions(-) create mode 100644 Mathlib/Topology/Algebra/Group/Quotient.lean diff --git a/Mathlib.lean b/Mathlib.lean index 5c43f1b7e519b..1ecdac31212ba 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4499,6 +4499,7 @@ import Mathlib.Topology.Algebra.FilterBasis import Mathlib.Topology.Algebra.Group.Basic import Mathlib.Topology.Algebra.Group.Compact import Mathlib.Topology.Algebra.Group.OpenMapping +import Mathlib.Topology.Algebra.Group.Quotient import Mathlib.Topology.Algebra.Group.SubmonoidClosure import Mathlib.Topology.Algebra.Group.TopologicalAbelianization import Mathlib.Topology.Algebra.GroupCompletion diff --git a/Mathlib/MeasureTheory/Constructions/Polish/Basic.lean b/Mathlib/MeasureTheory/Constructions/Polish/Basic.lean index 721c0f2411c57..02c0d1b8c30a6 100644 --- a/Mathlib/MeasureTheory/Constructions/Polish/Basic.lean +++ b/Mathlib/MeasureTheory/Constructions/Polish/Basic.lean @@ -631,10 +631,7 @@ instance CosetSpace.borelSpace {G : Type*} [TopologicalSpace G] [PolishSpace G] instance QuotientGroup.borelSpace {G : Type*} [TopologicalSpace G] [PolishSpace G] [Group G] [TopologicalGroup G] [MeasurableSpace G] [BorelSpace G] {N : Subgroup G} [N.Normal] [IsClosed (N : Set G)] : BorelSpace (G ⧸ N) := - -- Porting note: 1st and 3rd `haveI`s were not needed in Lean 3 - haveI := Subgroup.t3_quotient_of_isClosed N - haveI := QuotientGroup.secondCountableTopology (Γ := N) - Quotient.borelSpace + ⟨continuous_mk.map_eq_borel mk_surjective⟩ namespace MeasureTheory diff --git a/Mathlib/MeasureTheory/Measure/Haar/Quotient.lean b/Mathlib/MeasureTheory/Measure/Haar/Quotient.lean index 6cb8d92d7d95f..e552fa14c3286 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/Quotient.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/Quotient.lean @@ -57,7 +57,7 @@ instance QuotientGroup.measurableSMul {G : Type*} [Group G] {Γ : Subgroup G} [M [TopologicalSpace G] [TopologicalGroup G] [BorelSpace G] [BorelSpace (G ⧸ Γ)] : MeasurableSMul G (G ⧸ Γ) where measurable_const_smul g := (continuous_const_smul g).measurable - measurable_smul_const x := (QuotientGroup.continuous_smul₁ x).measurable + measurable_smul_const _ := (continuous_id.smul continuous_const).measurable end diff --git a/Mathlib/Topology/Algebra/Group/Basic.lean b/Mathlib/Topology/Algebra/Group/Basic.lean index 6ea401ab9a949..bb19df20f7401 100644 --- a/Mathlib/Topology/Algebra/Group/Basic.lean +++ b/Mathlib/Topology/Algebra/Group/Basic.lean @@ -3,13 +3,9 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro, Patrick Massot -/ -import Mathlib.GroupTheory.GroupAction.ConjAct -import Mathlib.GroupTheory.GroupAction.Quotient -import Mathlib.Topology.Algebra.Monoid -import Mathlib.Topology.Algebra.Constructions -import Mathlib.Topology.Maps.OpenQuotient +import Mathlib.Algebra.Group.Subgroup.Pointwise import Mathlib.Algebra.Order.Archimedean.Basic -import Mathlib.GroupTheory.QuotientGroup.Basic +import Mathlib.Topology.Algebra.Monoid /-! # Topological groups @@ -847,87 +843,6 @@ theorem TopologicalGroup.exists_antitone_basis_nhds_one [FirstCountableTopology end TopologicalGroup -namespace QuotientGroup - -variable [TopologicalSpace G] [Group G] - -@[to_additive] -instance instTopologicalSpace (N : Subgroup G) : TopologicalSpace (G ⧸ N) := - instTopologicalSpaceQuotient - -@[to_additive] -instance [CompactSpace G] (N : Subgroup G) : CompactSpace (G ⧸ N) := - Quotient.compactSpace - -@[to_additive] -theorem quotientMap_mk (N : Subgroup G) : QuotientMap (mk : G → G ⧸ N) := - quotientMap_quot_mk - -@[to_additive] -theorem continuous_mk {N : Subgroup G} : Continuous (mk : G → G ⧸ N) := - continuous_quot_mk - -section ContinuousMul - -variable [ContinuousMul G] {N : Subgroup G} - -@[to_additive] -theorem isOpenMap_coe : IsOpenMap ((↑) : G → G ⧸ N) := isOpenMap_quotient_mk'_mul - -@[to_additive] -theorem isOpenQuotientMap_mk : IsOpenQuotientMap (mk : G → G ⧸ N) := - MulAction.isOpenQuotientMap_quotientMk - -@[to_additive (attr := simp)] -theorem dense_preimage_mk {s : Set (G ⧸ N)} : Dense ((↑) ⁻¹' s : Set G) ↔ Dense s := - isOpenQuotientMap_mk.dense_preimage_iff - -@[to_additive] -theorem dense_image_mk {s : Set G} : - Dense (mk '' s : Set (G ⧸ N)) ↔ Dense (s * (N : Set G)) := by - rw [← dense_preimage_mk, preimage_image_mk_eq_mul] - -@[to_additive] -instance instContinuousSMul : ContinuousSMul G (G ⧸ N) where - continuous_smul := by - rw [← (IsOpenQuotientMap.id.prodMap isOpenQuotientMap_mk).continuous_comp_iff] - exact continuous_mk.comp continuous_mul - -variable (N) - -/-- Neighborhoods in the quotient are precisely the map of neighborhoods in the prequotient. -/ -@[to_additive - "Neighborhoods in the quotient are precisely the map of neighborhoods in the prequotient."] -theorem nhds_eq (x : G) : 𝓝 (x : G ⧸ N) = Filter.map (↑) (𝓝 x) := - (isOpenQuotientMap_mk.map_nhds_eq _).symm - -@[to_additive] -instance instFirstCountableTopology [FirstCountableTopology G] : - FirstCountableTopology (G ⧸ N) where - nhds_generated_countable := mk_surjective.forall.2 fun x ↦ nhds_eq N x ▸ inferInstance - -@[to_additive (attr := deprecated (since := "2024-08-05"))] -theorem nhds_one_isCountablyGenerated [FirstCountableTopology G] [N.Normal] : - (𝓝 (1 : G ⧸ N)).IsCountablyGenerated := - inferInstance - -end ContinuousMul - -variable [TopologicalGroup G] (N : Subgroup G) - -@[to_additive] -instance instTopologicalGroup [N.Normal] : TopologicalGroup (G ⧸ N) where - continuous_mul := by - rw [← (isOpenQuotientMap_mk.prodMap isOpenQuotientMap_mk).continuous_comp_iff] - exact continuous_mk.comp continuous_mul - continuous_inv := continuous_inv.quotient_map' _ - -@[to_additive (attr := deprecated (since := "2024-08-05"))] -theorem _root_.topologicalGroup_quotient [N.Normal] : TopologicalGroup (G ⧸ N) := - instTopologicalGroup N - -end QuotientGroup - /-- A typeclass saying that `p : G × G ↦ p.1 - p.2` is a continuous function. This property automatically holds for topological additive groups but it also holds, e.g., for `ℝ≥0`. -/ class ContinuousSub (G : Type*) [TopologicalSpace G] [Sub G] : Prop where @@ -1250,13 +1165,6 @@ theorem IsClosed.mul_right_of_isCompact (ht : IsClosed t) (hs : IsCompact s) : rw [← image_op_smul] exact IsClosed.smul_left_of_isCompact ht (hs.image continuous_op) -@[to_additive] -theorem QuotientGroup.isClosedMap_coe {H : Subgroup G} (hH : IsCompact (H : Set G)) : - IsClosedMap ((↑) : G → G ⧸ H) := by - intro t ht - rw [← (quotientMap_mk H).isClosed_preimage, preimage_image_mk_eq_mul] - exact ht.mul_right_of_isCompact hH - @[to_additive] lemma subset_mul_closure_one {G} [MulOneClass G] [TopologicalSpace G] (s : Set G) : s ⊆ s * (closure {1} : Set G) := by @@ -1382,13 +1290,6 @@ theorem exists_closed_nhds_one_inv_eq_mul_subset {U : Set G} (hU : U ∈ 𝓝 1) variable (S : Subgroup G) [Subgroup.Normal S] [IsClosed (S : Set G)] -@[to_additive] -instance Subgroup.t3_quotient_of_isClosed (S : Subgroup G) [Subgroup.Normal S] - [hS : IsClosed (S : Set G)] : T3Space (G ⧸ S) := by - rw [← QuotientGroup.ker_mk' S] at hS - haveI := TopologicalGroup.t1Space (G ⧸ S) (quotientMap_quotient_mk'.isClosed_preimage.mp hS) - infer_instance - /-- A subgroup `S` of a topological group `G` acts on `G` properly discontinuously on the left, if it is discrete in the sense that `S ∩ K` is finite for all compact `K`. (See also `DiscreteTopology`.) -/ @@ -1618,18 +1519,6 @@ theorem exists_isCompact_isClosed_nhds_one [WeaklyLocallyCompactSpace G] : let ⟨K, hK₁, hKcomp, hKcl⟩ := exists_mem_nhds_isCompact_isClosed (1 : G) ⟨K, hKcomp, hKcl, hK₁⟩ -/-- A quotient of a locally compact group is locally compact. -/ -@[to_additive] -instance [LocallyCompactSpace G] (N : Subgroup G) : LocallyCompactSpace (G ⧸ N) := by - refine ⟨fun x n hn ↦ ?_⟩ - let π := ((↑) : G → G ⧸ N) - have C : Continuous π := continuous_quotient_mk' - obtain ⟨y, rfl⟩ : ∃ y, π y = x := Quot.exists_rep x - have : π ⁻¹' n ∈ 𝓝 y := preimage_nhds_coinduced hn - rcases local_compact_nhds this with ⟨s, s_mem, hs, s_comp⟩ - exact ⟨π '' s, QuotientGroup.isOpenMap_coe.image_mem_nhds s_mem, mapsTo'.mp hs, - s_comp.image C⟩ - end section @@ -1665,30 +1554,6 @@ instance {G} [TopologicalSpace G] [AddGroup G] [TopologicalAddGroup G] : TopologicalGroup (Multiplicative G) where continuous_inv := @continuous_neg G _ _ _ -section Quotient - -variable [Group G] [TopologicalSpace G] [ContinuousMul G] {Γ : Subgroup G} - -@[to_additive] -instance QuotientGroup.continuousConstSMul : ContinuousConstSMul G (G ⧸ Γ) where - continuous_const_smul g := by - convert ((@continuous_const _ _ _ _ g).mul continuous_id).quotient_map' _ - -@[to_additive] -theorem QuotientGroup.continuous_smul₁ (x : G ⧸ Γ) : Continuous fun g : G => g • x := by - induction x using QuotientGroup.induction_on - exact continuous_quotient_mk'.comp (continuous_mul_right _) - -/-- The quotient of a second countable topological group by a subgroup is second countable. -/ -@[to_additive - "The quotient of a second countable additive topological group by a subgroup is second - countable."] -instance QuotientGroup.secondCountableTopology [SecondCountableTopology G] : - SecondCountableTopology (G ⧸ Γ) := - ContinuousConstSMul.secondCountableTopology - -end Quotient - /-- If `G` is a group with topological `⁻¹`, then it is homeomorphic to its units. -/ @[to_additive " If `G` is an additive group with topological negation, then it is homeomorphic to its additive units."] @@ -1944,4 +1809,4 @@ theorem coinduced_continuous {α β : Type*} [t : TopologicalSpace α] [Group β end GroupTopology -set_option linter.style.longFile 2100 +set_option linter.style.longFile 2000 diff --git a/Mathlib/Topology/Algebra/Group/Quotient.lean b/Mathlib/Topology/Algebra/Group/Quotient.lean new file mode 100644 index 0000000000000..08cb4e12c8933 --- /dev/null +++ b/Mathlib/Topology/Algebra/Group/Quotient.lean @@ -0,0 +1,133 @@ +/- +Copyright (c) 2017 Johannes Hölzl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johannes Hölzl, Mario Carneiro, Patrick Massot, Yury Kudryashov +-/ +import Mathlib.GroupTheory.GroupAction.Quotient +import Mathlib.GroupTheory.QuotientGroup.Basic +import Mathlib.Topology.Algebra.Group.Basic +import Mathlib.Topology.Maps.OpenQuotient + +/-! +# Topology on the quotient group + +In this file we define topology on `G ⧸ N`, where `N` is a subgroup of `G`, +and prove basic properties of this topology. +-/ + +open scoped Pointwise Topology + +variable {G : Type*} [TopologicalSpace G] [Group G] + +namespace QuotientGroup + +@[to_additive] +instance instTopologicalSpace (N : Subgroup G) : TopologicalSpace (G ⧸ N) := + instTopologicalSpaceQuotient + +@[to_additive] +instance [CompactSpace G] (N : Subgroup G) : CompactSpace (G ⧸ N) := + Quotient.compactSpace + +@[to_additive] +theorem quotientMap_mk (N : Subgroup G) : QuotientMap (mk : G → G ⧸ N) := + quotientMap_quot_mk + +@[to_additive] +theorem continuous_mk {N : Subgroup G} : Continuous (mk : G → G ⧸ N) := + continuous_quot_mk + +section ContinuousMul + +variable [ContinuousMul G] {N : Subgroup G} + +@[to_additive] +theorem isOpenMap_coe : IsOpenMap ((↑) : G → G ⧸ N) := isOpenMap_quotient_mk'_mul + +@[to_additive] +theorem isOpenQuotientMap_mk : IsOpenQuotientMap (mk : G → G ⧸ N) := + MulAction.isOpenQuotientMap_quotientMk + +@[to_additive (attr := simp)] +theorem dense_preimage_mk {s : Set (G ⧸ N)} : Dense ((↑) ⁻¹' s : Set G) ↔ Dense s := + isOpenQuotientMap_mk.dense_preimage_iff + +@[to_additive] +theorem dense_image_mk {s : Set G} : + Dense (mk '' s : Set (G ⧸ N)) ↔ Dense (s * (N : Set G)) := by + rw [← dense_preimage_mk, preimage_image_mk_eq_mul] + +@[to_additive] +instance instContinuousSMul : ContinuousSMul G (G ⧸ N) where + continuous_smul := by + rw [← (IsOpenQuotientMap.id.prodMap isOpenQuotientMap_mk).continuous_comp_iff] + exact continuous_mk.comp continuous_mul + +@[to_additive] +instance instContinuousConstSMul : ContinuousConstSMul G (G ⧸ N) := inferInstance + +/-- A quotient of a locally compact group is locally compact. -/ +@[to_additive] +instance instLocallyCompactSpace [LocallyCompactSpace G] (N : Subgroup G) : + LocallyCompactSpace (G ⧸ N) := + QuotientGroup.isOpenQuotientMap_mk.locallyCompactSpace + +@[to_additive (attr := deprecated (since := "2024-10-05"))] +theorem continuous_smul₁ (x : G ⧸ N) : Continuous fun g : G => g • x := + continuous_id.smul continuous_const + +variable (N) + +/-- Neighborhoods in the quotient are precisely the map of neighborhoods in the prequotient. -/ +@[to_additive + "Neighborhoods in the quotient are precisely the map of neighborhoods in the prequotient."] +theorem nhds_eq (x : G) : 𝓝 (x : G ⧸ N) = Filter.map (↑) (𝓝 x) := + (isOpenQuotientMap_mk.map_nhds_eq _).symm + +@[to_additive] +instance instFirstCountableTopology [FirstCountableTopology G] : + FirstCountableTopology (G ⧸ N) where + nhds_generated_countable := mk_surjective.forall.2 fun x ↦ nhds_eq N x ▸ inferInstance + +/-- The quotient of a second countable topological group by a subgroup is second countable. -/ +@[to_additive + "The quotient of a second countable additive topological group by a subgroup is second + countable."] +instance instSecondCountableTopology [SecondCountableTopology G] : + SecondCountableTopology (G ⧸ N) := + ContinuousConstSMul.secondCountableTopology + +@[to_additive (attr := deprecated (since := "2024-08-05"))] +theorem nhds_one_isCountablyGenerated [FirstCountableTopology G] [N.Normal] : + (𝓝 (1 : G ⧸ N)).IsCountablyGenerated := + inferInstance + +end ContinuousMul + +variable [TopologicalGroup G] (N : Subgroup G) + +@[to_additive] +instance instTopologicalGroup [N.Normal] : TopologicalGroup (G ⧸ N) where + continuous_mul := by + rw [← (isOpenQuotientMap_mk.prodMap isOpenQuotientMap_mk).continuous_comp_iff] + exact continuous_mk.comp continuous_mul + continuous_inv := continuous_inv.quotient_map' _ + +@[to_additive (attr := deprecated (since := "2024-08-05"))] +theorem _root_.topologicalGroup_quotient [N.Normal] : TopologicalGroup (G ⧸ N) := + instTopologicalGroup N + +@[to_additive] +theorem isClosedMap_coe {H : Subgroup G} (hH : IsCompact (H : Set G)) : + IsClosedMap ((↑) : G → G ⧸ H) := by + intro t ht + rw [← (quotientMap_mk H).isClosed_preimage, preimage_image_mk_eq_mul] + exact ht.mul_right_of_isCompact hH + +@[to_additive] +instance instT3Space [N.Normal] [hN : IsClosed (N : Set G)] : T3Space (G ⧸ N) := by + rw [← QuotientGroup.ker_mk' N] at hN + haveI := TopologicalGroup.t1Space (G ⧸ N) ((quotientMap_mk N).isClosed_preimage.mp hN) + infer_instance + +end QuotientGroup diff --git a/Mathlib/Topology/Algebra/Module/Basic.lean b/Mathlib/Topology/Algebra/Module/Basic.lean index c5fd7f2684f75..1ed9c07b215fe 100644 --- a/Mathlib/Topology/Algebra/Module/Basic.lean +++ b/Mathlib/Topology/Algebra/Module/Basic.lean @@ -2393,7 +2393,7 @@ instance continuousSMul_quotient [TopologicalSpace R] [TopologicalAddGroup M] [C instance t3_quotient_of_isClosed [TopologicalAddGroup M] [IsClosed (S : Set M)] : T3Space (M ⧸ S) := letI : IsClosed (S.toAddSubgroup : Set M) := ‹_› - S.toAddSubgroup.t3_quotient_of_isClosed + QuotientAddGroup.instT3Space S.toAddSubgroup end Submodule diff --git a/Mathlib/Topology/Algebra/OpenSubgroup.lean b/Mathlib/Topology/Algebra/OpenSubgroup.lean index 99c3a4c276f10..843e91673869b 100644 --- a/Mathlib/Topology/Algebra/OpenSubgroup.lean +++ b/Mathlib/Topology/Algebra/OpenSubgroup.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin, Nailin Guan -/ import Mathlib.RingTheory.Ideal.Basic +import Mathlib.Topology.Algebra.Group.Quotient import Mathlib.Topology.Algebra.Ring.Basic import Mathlib.Topology.Sets.Opens diff --git a/Mathlib/Topology/Algebra/ProperAction/Basic.lean b/Mathlib/Topology/Algebra/ProperAction/Basic.lean index 55f73de488d07..fa7a99510f216 100644 --- a/Mathlib/Topology/Algebra/ProperAction/Basic.lean +++ b/Mathlib/Topology/Algebra/ProperAction/Basic.lean @@ -5,6 +5,7 @@ Authors: Anatole Dedeker, Etienne Marion, Florestan Martin-Baillon, Vincent Guir -/ import Mathlib.Topology.Algebra.MulAction import Mathlib.Topology.Maps.Proper.Basic +import Mathlib.Topology.Maps.OpenQuotient /-! # Proper group action diff --git a/Mathlib/Topology/Algebra/Ring/Ideal.lean b/Mathlib/Topology/Algebra/Ring/Ideal.lean index aa3cf06eb57a1..c72e3ee7afefa 100644 --- a/Mathlib/Topology/Algebra/Ring/Ideal.lean +++ b/Mathlib/Topology/Algebra/Ring/Ideal.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Patrick Massot -/ import Mathlib.Topology.Algebra.Ring.Basic +import Mathlib.Topology.Algebra.Group.Quotient import Mathlib.RingTheory.Ideal.Quotient /-! diff --git a/Mathlib/Topology/Algebra/UniformGroup.lean b/Mathlib/Topology/Algebra/UniformGroup.lean index c05bd7340d340..1c3bc90dfda28 100644 --- a/Mathlib/Topology/Algebra/UniformGroup.lean +++ b/Mathlib/Topology/Algebra/UniformGroup.lean @@ -7,7 +7,7 @@ import Mathlib.Topology.UniformSpace.UniformConvergence import Mathlib.Topology.UniformSpace.UniformEmbedding import Mathlib.Topology.UniformSpace.CompleteSeparated import Mathlib.Topology.UniformSpace.Compact -import Mathlib.Topology.Algebra.Group.Basic +import Mathlib.Topology.Algebra.Group.Quotient import Mathlib.Topology.DiscreteSubset import Mathlib.Tactic.Abel diff --git a/Mathlib/Topology/Compactness/LocallyCompact.lean b/Mathlib/Topology/Compactness/LocallyCompact.lean index e0c44089ecd15..8d10e1ce3da2e 100644 --- a/Mathlib/Topology/Compactness/LocallyCompact.lean +++ b/Mathlib/Topology/Compactness/LocallyCompact.lean @@ -45,6 +45,13 @@ protected theorem IsClosed.weaklyLocallyCompactSpace [WeaklyLocallyCompactSpace {s : Set X} (hs : IsClosed s) : WeaklyLocallyCompactSpace s := (closedEmbedding_subtype_val hs).weaklyLocallyCompactSpace +theorem IsOpenQuotientMap.weaklyLocallyCompactSpace [WeaklyLocallyCompactSpace X] + {f : X → Y} (hf : IsOpenQuotientMap f) : WeaklyLocallyCompactSpace Y where + exists_compact_mem_nhds := by + refine hf.surjective.forall.2 fun x ↦ ?_ + rcases exists_compact_mem_nhds x with ⟨K, hKc, hKx⟩ + exact ⟨f '' K, hKc.image hf.continuous, hf.isOpenMap.image_mem_nhds hKx⟩ + /-- In a weakly locally compact space, every compact set is contained in the interior of a compact set. -/ theorem exists_compact_superset [WeaklyLocallyCompactSpace X] {K : Set X} (hK : IsCompact K) : @@ -166,6 +173,13 @@ theorem exists_compact_between [LocallyCompactSpace X] {K U : Set X} (hK : IsCom let ⟨L, hKL, hL, hLU⟩ := exists_mem_nhdsSet_isCompact_mapsTo continuous_id hK hU h_KU ⟨L, hL, subset_interior_iff_mem_nhdsSet.2 hKL, hLU⟩ +theorem IsOpenQuotientMap.locallyCompactSpace [LocallyCompactSpace X] {f : X → Y} + (hf : IsOpenQuotientMap f) : LocallyCompactSpace Y where + local_compact_nhds := by + refine hf.surjective.forall.2 fun x U hU ↦ ?_ + rcases local_compact_nhds (hf.continuous.continuousAt hU) with ⟨K, hKx, hKU, hKc⟩ + exact ⟨f '' K, hf.isOpenMap.image_mem_nhds hKx, image_subset_iff.2 hKU, hKc.image hf.continuous⟩ + /-- If `f` is a topology inducing map with a locally compact codomain and a locally closed range, then the domain of `f` is a locally compact space. -/ theorem Inducing.locallyCompactSpace [LocallyCompactSpace Y] {f : X → Y} (hf : Inducing f) From b1267a6b11cd3e2d499cd090a22e38adfed751e0 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 17 Oct 2024 03:45:31 +0000 Subject: [PATCH 043/505] feat: sigma-algebra of cylinder events (#17833) From GibbsMeasure Co-authored-by: Kin Yau James Wong --- .../Constructions/Cylinders.lean | 81 ++++++++++++++++++- 1 file changed, 79 insertions(+), 2 deletions(-) diff --git a/Mathlib/MeasureTheory/Constructions/Cylinders.lean b/Mathlib/MeasureTheory/Constructions/Cylinders.lean index 3ef19d096a73e..1f30b5fc6437d 100644 --- a/Mathlib/MeasureTheory/Constructions/Cylinders.lean +++ b/Mathlib/MeasureTheory/Constructions/Cylinders.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2023 Rémy Degenne. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Rémy Degenne, Peter Pfaffelhuber +Authors: Rémy Degenne, Peter Pfaffelhuber, Yaël Dillies, Kin Yau James Wong -/ import Mathlib.MeasureTheory.PiSystem import Mathlib.Order.OmegaCompletePartialOrder @@ -29,6 +29,8 @@ a product set. all `i` in the finset defining the box, the projection to `α i` belongs to `C i`. The main application of this is with `C i = {s : Set (α i) | MeasurableSet s}`. * `measurableCylinders`: set of all cylinders with measurable base sets. +* `cylinderEvents Δ`: The σ-algebra of cylinder events on `Δ`. It is the smallest σ-algebra making + the projections on the `i`-th coordinate continuous for all `i ∈ Δ`. ## Main statements @@ -39,7 +41,7 @@ a product set. -/ -open Set +open Function Set namespace MeasureTheory @@ -361,4 +363,79 @@ theorem generateFrom_measurableCylinders : end cylinders +/-! ### Cylinder events as a sigma-algebra -/ + +section cylinderEvents + +variable {α ι : Type*} {π : ι → Type*} {mα : MeasurableSpace α} [m : ∀ i, MeasurableSpace (π i)] + {Δ Δ₁ Δ₂ : Set ι} {i : ι} + +/-- The σ-algebra of cylinder events on `Δ`. It is the smallest σ-algebra making the projections +on the `i`-th coordinate continuous for all `i ∈ Δ`. -/ +def cylinderEvents (Δ : Set ι) : MeasurableSpace (∀ i, π i) := ⨆ i ∈ Δ, (m i).comap fun σ ↦ σ i + +@[simp] lemma cylinderEvents_univ : cylinderEvents (π := π) univ = MeasurableSpace.pi := by + simp [cylinderEvents, MeasurableSpace.pi] + +@[gcongr] +lemma cylinderEvents_mono (h : Δ₁ ⊆ Δ₂) : cylinderEvents (π := π) Δ₁ ≤ cylinderEvents Δ₂ := + biSup_mono h + +lemma cylinderEvents_le_pi : cylinderEvents (π := π) Δ ≤ MeasurableSpace.pi := by + simpa using cylinderEvents_mono (subset_univ _) + +lemma measurable_cylinderEvents_iff {g : α → ∀ i, π i} : + @Measurable _ _ _ (cylinderEvents Δ) g ↔ ∀ ⦃i⦄, i ∈ Δ → Measurable fun a ↦ g a i := by + simp_rw [measurable_iff_comap_le, cylinderEvents, MeasurableSpace.comap_iSup, + MeasurableSpace.comap_comp, Function.comp_def, iSup_le_iff] + +@[fun_prop, aesop safe 100 apply (rule_sets := [Measurable])] +lemma measurable_cylinderEvent_apply (hi : i ∈ Δ) : + Measurable[cylinderEvents Δ] fun f : ∀ i, π i => f i := + measurable_cylinderEvents_iff.1 measurable_id hi + +@[aesop safe 100 apply (rule_sets := [Measurable])] +lemma Measurable.eval_cylinderEvents {g : α → ∀ i, π i} (hi : i ∈ Δ) + (hg : @Measurable _ _ _ (cylinderEvents Δ) g) : Measurable fun a ↦ g a i := + (measurable_cylinderEvent_apply hi).comp hg + +@[fun_prop, aesop safe 100 apply (rule_sets := [Measurable])] +lemma measurable_cylinderEvents_lambda (f : α → ∀ i, π i) (hf : ∀ i, Measurable fun a ↦ f a i) : + Measurable f := + measurable_pi_iff.mpr hf + +/-- The function `(f, x) ↦ update f a x : (Π a, π a) × π a → Π a, π a` is measurable. -/ +lemma measurable_update_cylinderEvents' [DecidableEq ι] : + @Measurable _ _ (.prod (cylinderEvents Δ) (m i)) (cylinderEvents Δ) + (fun p : (∀ i, π i) × π i ↦ update p.1 i p.2) := by + rw [measurable_cylinderEvents_iff] + intro j hj + dsimp [update] + split_ifs with h + · subst h + dsimp + exact measurable_snd + · exact measurable_cylinderEvents_iff.1 measurable_fst hj + +lemma measurable_uniqueElim_cylinderEvents [Unique ι] : + Measurable (uniqueElim : π (default : ι) → ∀ i, π i) := by + simp_rw [measurable_pi_iff, Unique.forall_iff, uniqueElim_default]; exact measurable_id + +/-- The function `update f a : π a → Π a, π a` is always measurable. +This doesn't require `f` to be measurable. +This should not be confused with the statement that `update f a x` is measurable. -/ +@[measurability] +lemma measurable_update_cylinderEvents (f : ∀ a : ι, π a) {a : ι} [DecidableEq ι] : + @Measurable _ _ _ (cylinderEvents Δ) (update f a) := + measurable_update_cylinderEvents'.comp measurable_prod_mk_left + +lemma measurable_update_cylinderEvents_left {a : ι} [DecidableEq ι] {x : π a} : + @Measurable _ _ (cylinderEvents Δ) (cylinderEvents Δ) (update · a x) := + measurable_update_cylinderEvents'.comp measurable_prod_mk_right + +lemma measurable_restrict_cylinderEvents (Δ : Set ι) : + Measurable[cylinderEvents (π := π) Δ] (restrict Δ) := by + rw [@measurable_pi_iff]; exact fun i ↦ measurable_cylinderEvent_apply i.2 + +end cylinderEvents end MeasureTheory From 08804ac5cc5dca54e6979de8388f59c937b60e8a Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Thu, 17 Oct 2024 03:53:08 +0000 Subject: [PATCH 044/505] =?UTF-8?q?chore(SetTheory/Ordinal/Arithmetic):=20?= =?UTF-8?q?`omega0=5FisLimit`=20=E2=86=92=20`isLimit=5Fomega0`=20(#17744)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit I missed this rename in #17673. --- Mathlib/SetTheory/Cardinal/Aleph.lean | 6 +++--- Mathlib/SetTheory/Cardinal/Arithmetic.lean | 2 +- Mathlib/SetTheory/Cardinal/Cofinality.lean | 12 ++++++------ Mathlib/SetTheory/Ordinal/Arithmetic.lean | 20 +++++++++++++------- Mathlib/SetTheory/Ordinal/Notation.lean | 14 +++++++------- Mathlib/SetTheory/Ordinal/Principal.lean | 10 +++++----- 6 files changed, 35 insertions(+), 29 deletions(-) diff --git a/Mathlib/SetTheory/Cardinal/Aleph.lean b/Mathlib/SetTheory/Cardinal/Aleph.lean index 7a654d0363135..090dd44cb90c6 100644 --- a/Mathlib/SetTheory/Cardinal/Aleph.lean +++ b/Mathlib/SetTheory/Cardinal/Aleph.lean @@ -249,7 +249,7 @@ theorem aleph'_limit {o : Ordinal} (ho : o.IsLimit) : aleph' o = ⨆ a : Iio o, @[simp] theorem aleph'_omega0 : aleph' ω = ℵ₀ := eq_of_forall_ge_iff fun c => by - simp only [aleph'_le_of_limit omega0_isLimit, lt_omega0, exists_imp, aleph0_le] + simp only [aleph'_le_of_limit isLimit_omega0, lt_omega0, exists_imp, aleph0_le] exact forall_swap.trans (forall_congr' fun n => by simp only [forall_eq, aleph'_nat]) @[deprecated (since := "2024-09-30")] @@ -333,12 +333,12 @@ instance nonempty_toType_aleph (o : Ordinal) : Nonempty (ℵ_ o).ord.toType := b exact fun h => (ord_injective h).not_gt (aleph_pos o) theorem ord_aleph_isLimit (o : Ordinal) : (ℵ_ o).ord.IsLimit := - ord_isLimit <| aleph0_le_aleph _ + isLimit_ord <| aleph0_le_aleph _ instance (o : Ordinal) : NoMaxOrder (ℵ_ o).ord.toType := toType_noMax_of_succ_lt (ord_aleph_isLimit o).2 -theorem exists_aleph {c : Cardinal} : ℵ₀ ≤ c ↔ ∃ o, c = aleph o := +theorem exists_aleph {c : Cardinal} : ℵ₀ ≤ c ↔ ∃ o, c = ℵ_ o := ⟨fun h => ⟨aleph'.symm c - ω, by rw [aleph_eq_aleph', Ordinal.add_sub_cancel_of_le, aleph'.apply_symm_apply] diff --git a/Mathlib/SetTheory/Cardinal/Arithmetic.lean b/Mathlib/SetTheory/Cardinal/Arithmetic.lean index 4bb8721a053d7..4dfd399fb40da 100644 --- a/Mathlib/SetTheory/Cardinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Cardinal/Arithmetic.lean @@ -82,7 +82,7 @@ theorem mul_eq_self {c : Cardinal} (h : ℵ₀ ≤ c) : c * c = c := by · exact (mul_lt_aleph0 qo qo).trans_le ol · suffices (succ (typein LT.lt (g p))).card < ⟦α⟧ from (IH _ this qo).trans_lt this rw [← lt_ord] - apply (ord_isLimit ol).2 + apply (isLimit_ord ol).2 rw [mk'_def, e] apply typein_lt_type diff --git a/Mathlib/SetTheory/Cardinal/Cofinality.lean b/Mathlib/SetTheory/Cardinal/Cofinality.lean index 5f1d484af9be7..cbf48de84b877 100644 --- a/Mathlib/SetTheory/Cardinal/Cofinality.lean +++ b/Mathlib/SetTheory/Cardinal/Cofinality.lean @@ -681,7 +681,7 @@ theorem aleph_cof {o : Ordinal} (ho : o.IsLimit) : (aleph o).ord.cof = o.cof := @[simp] theorem cof_omega0 : cof ω = ℵ₀ := - (aleph0_le_cof.2 omega0_isLimit).antisymm' <| by + (aleph0_le_cof.2 isLimit_omega0).antisymm' <| by rw [← card_omega0] apply cof_le_card @@ -862,7 +862,7 @@ theorem mk_bounded_subset {α : Type*} (h : ∀ x < #α, (2^x) < #α) {r : α · refine @mk_le_of_injective α _ (fun x => Subtype.mk {x} ?_) ?_ · apply bounded_singleton rw [← hr] - apply ord_isLimit ha + apply isLimit_ord ha · intro a b hab simpa [singleton_eq_singleton_iff] using hab @@ -881,7 +881,7 @@ theorem mk_subset_mk_lt_cof {α : Type*} (h : ∀ x < #α, (2^x) < #α) : exact lt_cof_type hs · refine @mk_le_of_injective α _ (fun x => Subtype.mk {x} ?_) ?_ · rw [mk_singleton] - exact one_lt_aleph0.trans_le (aleph0_le_cof.2 (ord_isLimit h'.aleph0_le)) + exact one_lt_aleph0.trans_le (aleph0_le_cof.2 (isLimit_ord h'.aleph0_le)) · intro a b hab simpa [singleton_eq_singleton_iff] using hab @@ -917,7 +917,7 @@ theorem isRegular_succ {c : Cardinal.{u}} (h : ℵ₀ ≤ c) : IsRegular (succ c (by cases' Quotient.exists_rep (@succ Cardinal _ _ c) with α αe; simp only [mk'_def] at αe rcases ord_eq α with ⟨r, wo, re⟩ - have := ord_isLimit (h.trans (le_succ _)) + have := isLimit_ord (h.trans (le_succ _)) rw [← αe, re] at this ⊢ rcases cof_eq' r this with ⟨S, H, Se⟩ rw [← Se] @@ -1123,7 +1123,7 @@ theorem derivFamily_lt_ord_lift {ι} {f : ι → Ordinal → Ordinal} {c} (hc : rw [derivFamily_succ] exact nfpFamily_lt_ord_lift hω (by rwa [hc.cof_eq]) hf - ((ord_isLimit hc.1).2 _ (hb ((lt_succ b).trans hb'))) + ((isLimit_ord hc.1).2 _ (hb ((lt_succ b).trans hb'))) | H₃ b hb H => intro hb' rw [derivFamily_limit f hb] @@ -1170,7 +1170,7 @@ theorem univ_inaccessible : IsInaccessible univ.{u, v} := theorem lt_power_cof {c : Cardinal.{u}} : ℵ₀ ≤ c → c < (c^cof c.ord) := Quotient.inductionOn c fun α h => by rcases ord_eq α with ⟨r, wo, re⟩ - have := ord_isLimit h + have := isLimit_ord h rw [mk'_def, re] at this ⊢ rcases cof_eq' r this with ⟨S, H, Se⟩ have := sum_lt_prod (fun a : S => #{ x // r x a }) (fun _ => #α) fun i => ?_ diff --git a/Mathlib/SetTheory/Ordinal/Arithmetic.lean b/Mathlib/SetTheory/Ordinal/Arithmetic.lean index d8239b6b16acc..90d58666720d4 100644 --- a/Mathlib/SetTheory/Ordinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Ordinal/Arithmetic.lean @@ -2263,13 +2263,16 @@ theorem one_lt_omega0 : 1 < ω := by simpa only [Nat.cast_one] using nat_lt_omeg @[deprecated (since := "2024-09-30")] alias one_lt_omega := one_lt_omega0 -theorem omega0_isLimit : IsLimit ω := +theorem isLimit_omega0 : IsLimit ω := ⟨omega0_ne_zero, fun o h => by let ⟨n, e⟩ := lt_omega0.1 h rw [e]; exact nat_lt_omega0 (n + 1)⟩ +@[deprecated (since := "2024-10-14")] +alias omega0_isLimit := isLimit_omega0 + @[deprecated (since := "2024-09-30")] -alias omega_isLimit := omega0_isLimit +alias omega_isLimit := isLimit_omega0 theorem omega0_le {o : Ordinal} : ω ≤ o ↔ ∀ n : ℕ, ↑n ≤ o := ⟨fun h n => (nat_lt_omega0 _).le.trans h, fun H => @@ -2306,14 +2309,14 @@ theorem isLimit_iff_omega0_dvd {a : Ordinal} : IsLimit a ↔ a ≠ 0 ∧ ω ∣ refine ⟨fun l => ⟨l.1, ⟨a / ω, le_antisymm ?_ (mul_div_le _ _)⟩⟩, fun h => ?_⟩ · refine (limit_le l).2 fun x hx => le_of_lt ?_ rw [← div_lt omega0_ne_zero, ← succ_le_iff, le_div omega0_ne_zero, mul_succ, - add_le_of_limit omega0_isLimit] + add_le_of_limit isLimit_omega0] intro b hb rcases lt_omega0.1 hb with ⟨n, rfl⟩ exact (add_le_add_right (mul_div_le _ _) _).trans (lt_sub.1 <| nat_lt_limit (isLimit_sub l hx) _).le · rcases h with ⟨a0, b, rfl⟩ - refine isLimit_mul_left omega0_isLimit (Ordinal.pos_iff_ne_zero.2 <| mt ?_ a0) + refine isLimit_mul_left isLimit_omega0 (Ordinal.pos_iff_ne_zero.2 <| mt ?_ a0) intro e simp only [e, mul_zero] @@ -2394,7 +2397,7 @@ namespace Cardinal open Ordinal -theorem ord_isLimit {c} (co : ℵ₀ ≤ c) : (ord c).IsLimit := by +theorem isLimit_ord {c} (co : ℵ₀ ≤ c) : (ord c).IsLimit := by refine ⟨fun h => aleph0_ne_zero ?_, fun a => lt_imp_lt_of_le_imp_le fun h => ?_⟩ · rw [← Ordinal.le_zero, ord_le] at h simpa only [card_zero, nonpos_iff_eq_zero] using co.trans h @@ -2403,10 +2406,13 @@ theorem ord_isLimit {c} (co : ℵ₀ ≤ c) : (ord c).IsLimit := by rw [← ord_le, ← le_succ_of_isLimit, ord_le] · exact co.trans h · rw [ord_aleph0] - exact Ordinal.omega0_isLimit + exact Ordinal.isLimit_omega0 + +@[deprecated (since := "2024-10-14")] +alias ord_isLimit := isLimit_ord theorem noMaxOrder {c} (h : ℵ₀ ≤ c) : NoMaxOrder c.ord.toType := - toType_noMax_of_succ_lt (ord_isLimit h).2 + toType_noMax_of_succ_lt (isLimit_ord h).2 end Cardinal diff --git a/Mathlib/SetTheory/Ordinal/Notation.lean b/Mathlib/SetTheory/Ordinal/Notation.lean index 9df8635093a8b..050c6b4f8f653 100644 --- a/Mathlib/SetTheory/Ordinal/Notation.lean +++ b/Mathlib/SetTheory/Ordinal/Notation.lean @@ -577,7 +577,7 @@ theorem repr_mul : ∀ (o₁ o₂) [NF o₁] [NF o₂], repr (o₁ * o₂) = rep rw [← mul_assoc] congr 2 have := mt repr_inj.1 e0 - rw [add_mul_limit ao (isLimit_opow_left omega0_isLimit this), mul_assoc, + rw [add_mul_limit ao (isLimit_opow_left isLimit_omega0 this), mul_assoc, mul_omega0_dvd (natCast_pos.2 n₁.pos) (nat_lt_omega0 _)] simpa using opow_dvd_opow ω (one_le_iff_ne_zero.2 this) @@ -790,7 +790,7 @@ theorem repr_opow_aux₁ {e a} [Ne : NF e] [Na : NF a] {a' : Ordinal} (e0 : repr have := omega0_le_oadd e n a rw [repr] at this refine le_antisymm ?_ (opow_le_opow_left _ this) - apply (opow_le_of_limit ((opow_pos _ omega0_pos).trans_le this).ne' omega0_isLimit).2 + apply (opow_le_of_limit ((opow_pos _ omega0_pos).trans_le this).ne' isLimit_omega0).2 intro b l have := (No.below_of_lt (lt_succ _)).repr_lt rw [repr] at this @@ -801,8 +801,8 @@ theorem repr_opow_aux₁ {e a} [Ne : NF e] [Na : NF a] {a' : Ordinal} (e0 : repr · apply (mul_le_mul_left' (le_succ b) _).trans rw [← add_one_eq_succ, add_mul_succ _ (one_add_of_omega0_le h), add_one_eq_succ, succ_le_iff, Ordinal.mul_lt_mul_iff_left (Ordinal.pos_iff_ne_zero.2 e0)] - exact omega0_isLimit.2 _ l - · apply (principal_mul_omega0 (omega0_isLimit.2 _ h) l).le.trans + exact isLimit_omega0.2 _ l + · apply (principal_mul_omega0 (isLimit_omega0.2 _ h) l).le.trans simpa using mul_le_mul_right' (one_le_iff_ne_zero.2 e0) ω section @@ -966,7 +966,7 @@ private theorem exists_lt_add {α} [hα : Nonempty α] {o : Ordinal} {f : α → private theorem exists_lt_mul_omega0' {o : Ordinal} ⦃a⦄ (h : a < o * ω) : ∃ i : ℕ, a < o * ↑i + o := by - obtain ⟨i, hi, h'⟩ := (lt_mul_of_limit omega0_isLimit).1 h + obtain ⟨i, hi, h'⟩ := (lt_mul_of_limit isLimit_omega0).1 h obtain ⟨i, rfl⟩ := lt_omega0.1 hi exact ⟨i, h'.trans_le (le_add_right _ _)⟩ @@ -1025,13 +1025,13 @@ theorem fundamentalSequence_has_prop (o) : FundamentalSequenceProp o (fundamenta · exact ⟨rfl, inferInstance⟩ · have := opow_pos (repr a') omega0_pos refine - ⟨isLimit_mul this omega0_isLimit, fun i => + ⟨isLimit_mul this isLimit_omega0, fun i => ⟨this, ?_, fun H => @NF.oadd_zero _ _ (iha.2 H.fst)⟩, exists_lt_mul_omega0'⟩ rw [← mul_succ, ← natCast_succ, Ordinal.mul_lt_mul_iff_left this] apply nat_lt_omega0 · have := opow_pos (repr a') omega0_pos refine - ⟨isLimit_add _ (isLimit_mul this omega0_isLimit), fun i => ⟨this, ?_, ?_⟩, + ⟨isLimit_add _ (isLimit_mul this isLimit_omega0), fun i => ⟨this, ?_, ?_⟩, exists_lt_add exists_lt_mul_omega0'⟩ · rw [← mul_succ, ← natCast_succ, Ordinal.mul_lt_mul_iff_left this] apply nat_lt_omega0 diff --git a/Mathlib/SetTheory/Ordinal/Principal.lean b/Mathlib/SetTheory/Ordinal/Principal.lean index d58c0eff2dd35..5e32abd578e7d 100644 --- a/Mathlib/SetTheory/Ordinal/Principal.lean +++ b/Mathlib/SetTheory/Ordinal/Principal.lean @@ -194,7 +194,7 @@ theorem add_omega0_opow {a b : Ordinal} (h : a < ω ^ b) : a + ω ^ b = ω ^ b : · rw [opow_zero, ← succ_zero, lt_succ_iff, Ordinal.le_zero] at h rw [h, zero_add] · rw [opow_succ] at h - rcases (lt_mul_of_limit omega0_isLimit).1 h with ⟨x, xo, ax⟩ + rcases (lt_mul_of_limit isLimit_omega0).1 h with ⟨x, xo, ax⟩ apply (add_le_add_right ax.le _).trans rw [opow_succ, ← mul_add, add_omega0 xo] · rcases (lt_opow_of_limit omega0_ne_zero l).1 h with ⟨x, xb, ax⟩ @@ -228,7 +228,7 @@ theorem principal_add_iff_zero_or_omega0_opow {o : Ordinal} : fun ⟨b, e⟩ => e.symm ▸ fun a => add_omega0_opow⟩ have := H _ h have := lt_opow_succ_log_self one_lt_omega0 o - rw [opow_succ, lt_mul_of_limit omega0_isLimit] at this + rw [opow_succ, lt_mul_of_limit isLimit_omega0] at this rcases this with ⟨a, ao, h'⟩ rcases lt_omega0.1 ao with ⟨n, rfl⟩ clear ao @@ -347,7 +347,7 @@ theorem mul_lt_omega0_opow {a b c : Ordinal} (c0 : 0 < c) (ha : a < ω ^ c) (hb · exact (lt_irrefl _).elim c0 · rw [opow_succ] at ha obtain ⟨n, hn, an⟩ := - ((isNormal_mul_right <| opow_pos _ omega0_pos).limit_lt omega0_isLimit).1 ha + ((isNormal_mul_right <| opow_pos _ omega0_pos).limit_lt isLimit_omega0).1 ha apply (mul_le_mul_right' (le_of_lt an) _).trans_lt rw [opow_succ, mul_assoc, mul_lt_mul_iff_left (opow_pos _ omega0_pos)] exact principal_mul_omega0 hn hb @@ -366,7 +366,7 @@ theorem mul_omega0_opow_opow {a b : Ordinal} (a0 : 0 < a) (h : a < ω ^ ω ^ b) exact mul_omega0 a0 h · apply le_antisymm · obtain ⟨x, xb, ax⟩ := - (lt_opow_of_limit omega0_ne_zero (isLimit_opow_left omega0_isLimit b0)).1 h + (lt_opow_of_limit omega0_ne_zero (isLimit_opow_left isLimit_omega0 b0)).1 h apply (mul_le_mul_right' (le_of_lt ax) _).trans rw [← opow_add, add_omega0_opow xb] · conv_lhs => rw [← one_mul (ω ^ _)] @@ -442,7 +442,7 @@ theorem principal_opow_omega0 : Principal (· ^ ·) ω := fun a b ha hb => alias principal_opow_omega := principal_opow_omega0 theorem opow_omega0 {a : Ordinal} (a1 : 1 < a) (h : a < ω) : a ^ ω = ω := - ((opow_le_of_limit (one_le_iff_ne_zero.1 <| le_of_lt a1) omega0_isLimit).2 fun _ hb => + ((opow_le_of_limit (one_le_iff_ne_zero.1 <| le_of_lt a1) isLimit_omega0).2 fun _ hb => (principal_opow_omega0 h hb).le).antisymm (right_le_opow _ a1) From 8e52ebf89bc31f366f732f12e4ed932eeecf442e Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Thu, 17 Oct 2024 04:14:09 +0000 Subject: [PATCH 045/505] chore: move differentiability of Gamma to a new file (#17775) --- Mathlib.lean | 1 + Mathlib/Analysis/MellinInversion.lean | 1 + .../SpecialFunctions/Gamma/Basic.lean | 85 +------------ .../SpecialFunctions/Gamma/BohrMollerup.lean | 1 + .../SpecialFunctions/Gamma/Deriv.lean | 117 ++++++++++++++++++ 5 files changed, 122 insertions(+), 83 deletions(-) create mode 100644 Mathlib/Analysis/SpecialFunctions/Gamma/Deriv.lean diff --git a/Mathlib.lean b/Mathlib.lean index 1ecdac31212ba..14f4a5f1fa41b 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1347,6 +1347,7 @@ import Mathlib.Analysis.SpecialFunctions.Gamma.Basic import Mathlib.Analysis.SpecialFunctions.Gamma.Beta import Mathlib.Analysis.SpecialFunctions.Gamma.BohrMollerup import Mathlib.Analysis.SpecialFunctions.Gamma.Deligne +import Mathlib.Analysis.SpecialFunctions.Gamma.Deriv import Mathlib.Analysis.SpecialFunctions.Gaussian.FourierTransform import Mathlib.Analysis.SpecialFunctions.Gaussian.GaussianIntegral import Mathlib.Analysis.SpecialFunctions.Gaussian.PoissonSummation diff --git a/Mathlib/Analysis/MellinInversion.lean b/Mathlib/Analysis/MellinInversion.lean index c9680254f0421..758a80e2575a8 100644 --- a/Mathlib/Analysis/MellinInversion.lean +++ b/Mathlib/Analysis/MellinInversion.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Lawrence Wu -/ import Mathlib.Analysis.Fourier.Inversion +import Mathlib.Analysis.MellinTransform /-! # Mellin inversion formula diff --git a/Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean b/Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean index 685bc7386b0d0..8a5e577497057 100644 --- a/Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean +++ b/Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean @@ -3,8 +3,8 @@ Copyright (c) 2022 David Loeffler. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: David Loeffler -/ +import Mathlib.Analysis.SpecialFunctions.ImproperIntegrals import Mathlib.MeasureTheory.Integral.ExpDecay -import Mathlib.Analysis.MellinTransform /-! # The Gamma function @@ -24,15 +24,12 @@ set it to be `0` by convention.) * `Complex.Gamma_eq_integral`: for `0 < re s`, `Γ(s)` agrees with Euler's integral. * `Complex.Gamma_add_one`: for all `s : ℂ` with `s ≠ 0`, we have `Γ (s + 1) = s Γ(s)`. * `Complex.Gamma_nat_eq_factorial`: for all `n : ℕ` we have `Γ (n + 1) = n!`. -* `Complex.differentiableAt_Gamma`: `Γ` is complex-differentiable at all `s : ℂ` with - `s ∉ {-n : n ∈ ℕ}`. ## Gamma function: main statements (real case) * `Real.Gamma`: the `Γ` function (of a real variable). * Real counterparts of all the properties of the complex Gamma function listed above: - `Real.Gamma_eq_integral`, `Real.Gamma_add_one`, `Real.Gamma_nat_eq_factorial`, - `Real.differentiableAt_Gamma`. + `Real.Gamma_eq_integral`, `Real.Gamma_add_one`, `Real.Gamma_nat_eq_factorial`. ## Tags @@ -395,79 +392,6 @@ lemma integral_cpow_mul_exp_neg_mul_Ioi {a : ℂ} {r : ℝ} (ha : 0 < a.re) (hr end GammaDef -/-! Now check that the `Γ` function is differentiable, wherever this makes sense. -/ - - -section GammaHasDeriv - -/-- Rewrite the Gamma integral as an example of a Mellin transform. -/ -theorem GammaIntegral_eq_mellin : GammaIntegral = mellin fun x => ↑(Real.exp (-x)) := - funext fun s => by simp only [mellin, GammaIntegral, smul_eq_mul, mul_comm] - -/-- The derivative of the `Γ` integral, at any `s ∈ ℂ` with `1 < re s`, is given by the Mellin -transform of `log t * exp (-t)`. -/ -theorem hasDerivAt_GammaIntegral {s : ℂ} (hs : 0 < s.re) : - HasDerivAt GammaIntegral (∫ t : ℝ in Ioi 0, t ^ (s - 1) * (Real.log t * Real.exp (-t))) s := by - rw [GammaIntegral_eq_mellin] - convert (mellin_hasDerivAt_of_isBigO_rpow (E := ℂ) _ _ (lt_add_one _) _ hs).2 - · refine (Continuous.continuousOn ?_).locallyIntegrableOn measurableSet_Ioi - exact continuous_ofReal.comp (Real.continuous_exp.comp continuous_neg) - · rw [← isBigO_norm_left] - simp_rw [Complex.norm_eq_abs, abs_ofReal, ← Real.norm_eq_abs, isBigO_norm_left] - simpa only [neg_one_mul] using (isLittleO_exp_neg_mul_rpow_atTop zero_lt_one _).isBigO - · simp_rw [neg_zero, rpow_zero] - refine isBigO_const_of_tendsto (?_ : Tendsto _ _ (𝓝 (1 : ℂ))) one_ne_zero - rw [(by simp : (1 : ℂ) = Real.exp (-0))] - exact (continuous_ofReal.comp (Real.continuous_exp.comp continuous_neg)).continuousWithinAt - -theorem differentiableAt_GammaAux (s : ℂ) (n : ℕ) (h1 : 1 - s.re < n) (h2 : ∀ m : ℕ, s ≠ -m) : - DifferentiableAt ℂ (GammaAux n) s := by - induction' n with n hn generalizing s - · refine (hasDerivAt_GammaIntegral ?_).differentiableAt - rw [Nat.cast_zero] at h1; linarith - · dsimp only [GammaAux] - specialize hn (s + 1) - have a : 1 - (s + 1).re < ↑n := by - rw [Nat.cast_succ] at h1; rw [Complex.add_re, Complex.one_re]; linarith - have b : ∀ m : ℕ, s + 1 ≠ -m := by - intro m; have := h2 (1 + m) - contrapose! this - rw [← eq_sub_iff_add_eq] at this - simpa using this - refine DifferentiableAt.div (DifferentiableAt.comp _ (hn a b) ?_) ?_ ?_ - · rw [differentiableAt_add_const_iff (1 : ℂ)]; exact differentiableAt_id - · exact differentiableAt_id - · simpa using h2 0 - -theorem differentiableAt_Gamma (s : ℂ) (hs : ∀ m : ℕ, s ≠ -m) : DifferentiableAt ℂ Gamma s := by - let n := ⌊1 - s.re⌋₊ + 1 - have hn : 1 - s.re < n := mod_cast Nat.lt_floor_add_one (1 - s.re) - apply (differentiableAt_GammaAux s n hn hs).congr_of_eventuallyEq - let S := {t : ℂ | 1 - t.re < n} - have : S ∈ 𝓝 s := by - rw [mem_nhds_iff]; use S - refine ⟨Subset.rfl, ?_, hn⟩ - have : S = re ⁻¹' Ioi (1 - n : ℝ) := by - ext; rw [preimage, Ioi, mem_setOf_eq, mem_setOf_eq, mem_setOf_eq]; exact sub_lt_comm - rw [this] - exact Continuous.isOpen_preimage continuous_re _ isOpen_Ioi - apply eventuallyEq_of_mem this - intro t ht; rw [mem_setOf_eq] at ht - apply Gamma_eq_GammaAux; linarith - -end GammaHasDeriv - -/-- At `s = 0`, the Gamma function has a simple pole with residue 1. -/ -theorem tendsto_self_mul_Gamma_nhds_zero : Tendsto (fun z : ℂ => z * Gamma z) (𝓝[≠] 0) (𝓝 1) := by - rw [show 𝓝 (1 : ℂ) = 𝓝 (Gamma (0 + 1)) by simp only [zero_add, Complex.Gamma_one]] - convert (Tendsto.mono_left _ nhdsWithin_le_nhds).congr' - (eventuallyEq_of_mem self_mem_nhdsWithin Complex.Gamma_add_one) using 1 - refine ContinuousAt.comp (g := Gamma) ?_ (continuous_id.add continuous_const).continuousAt - refine (Complex.differentiableAt_Gamma _ fun m => ?_).continuousAt - rw [zero_add, ← ofReal_natCast, ← ofReal_neg, ← ofReal_one, Ne, ofReal_inj] - refine (lt_of_le_of_lt ?_ zero_lt_one).ne' - exact neg_nonpos.mpr (Nat.cast_nonneg _) - end Complex namespace Real @@ -606,9 +530,4 @@ theorem Gamma_ne_zero {s : ℝ} (hs : ∀ m : ℕ, s ≠ -m) : Gamma s ≠ 0 := theorem Gamma_eq_zero_iff (s : ℝ) : Gamma s = 0 ↔ ∃ m : ℕ, s = -m := ⟨by contrapose!; exact Gamma_ne_zero, by rintro ⟨m, rfl⟩; exact Gamma_neg_nat_eq_zero m⟩ -theorem differentiableAt_Gamma {s : ℝ} (hs : ∀ m : ℕ, s ≠ -m) : DifferentiableAt ℝ Gamma s := by - refine (Complex.differentiableAt_Gamma _ ?_).hasDerivAt.real_of_complex.differentiableAt - simp_rw [← Complex.ofReal_natCast, ← Complex.ofReal_neg, Ne, Complex.ofReal_inj] - exact hs - end Real diff --git a/Mathlib/Analysis/SpecialFunctions/Gamma/BohrMollerup.lean b/Mathlib/Analysis/SpecialFunctions/Gamma/BohrMollerup.lean index c9f1f89b0bd5a..f777ccb89b1bb 100644 --- a/Mathlib/Analysis/SpecialFunctions/Gamma/BohrMollerup.lean +++ b/Mathlib/Analysis/SpecialFunctions/Gamma/BohrMollerup.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 David Loeffler. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: David Loeffler -/ +import Mathlib.Analysis.SpecialFunctions.Gamma.Deriv import Mathlib.Analysis.SpecialFunctions.Gaussian.GaussianIntegral /-! # Convexity properties of the Gamma function diff --git a/Mathlib/Analysis/SpecialFunctions/Gamma/Deriv.lean b/Mathlib/Analysis/SpecialFunctions/Gamma/Deriv.lean new file mode 100644 index 0000000000000..8488967aba430 --- /dev/null +++ b/Mathlib/Analysis/SpecialFunctions/Gamma/Deriv.lean @@ -0,0 +1,117 @@ +/- +Copyright (c) 2022 David Loeffler. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: David Loeffler +-/ +import Mathlib.Analysis.MellinTransform +import Mathlib.Analysis.SpecialFunctions.Gamma.Basic + +/-! +# Derivative of the Gamma function + +This file shows that the (complex) `Γ` function is complex-differentiable at all `s : ℂ` with +`s ∉ {-n : n ∈ ℕ}`, as well as the real counterpart. + +## Main results + +* `Complex.differentiableAt_Gamma`: `Γ` is complex-differentiable at all `s : ℂ` with + `s ∉ {-n : n ∈ ℕ}`. +* `Real.differentiableAt_Gamma`: `Γ` is real-differentiable at all `s : ℝ` with + `s ∉ {-n : n ∈ ℕ}`. + +## Tags + +Gamma +-/ + + +noncomputable section + +open Filter Set Real Asymptotics +open scoped Topology + +namespace Complex + +/-! Now check that the `Γ` function is differentiable, wherever this makes sense. -/ + + +section GammaHasDeriv + +/-- Rewrite the Gamma integral as an example of a Mellin transform. -/ +theorem GammaIntegral_eq_mellin : GammaIntegral = mellin fun x => ↑(Real.exp (-x)) := + funext fun s => by simp only [mellin, GammaIntegral, smul_eq_mul, mul_comm] + +/-- The derivative of the `Γ` integral, at any `s ∈ ℂ` with `1 < re s`, is given by the Mellin +transform of `log t * exp (-t)`. -/ +theorem hasDerivAt_GammaIntegral {s : ℂ} (hs : 0 < s.re) : + HasDerivAt GammaIntegral (∫ t : ℝ in Ioi 0, t ^ (s - 1) * (Real.log t * Real.exp (-t))) s := by + rw [GammaIntegral_eq_mellin] + convert (mellin_hasDerivAt_of_isBigO_rpow (E := ℂ) _ _ (lt_add_one _) _ hs).2 + · refine (Continuous.continuousOn ?_).locallyIntegrableOn measurableSet_Ioi + exact continuous_ofReal.comp (Real.continuous_exp.comp continuous_neg) + · rw [← isBigO_norm_left] + simp_rw [Complex.norm_eq_abs, abs_ofReal, ← Real.norm_eq_abs, isBigO_norm_left] + simpa only [neg_one_mul] using (isLittleO_exp_neg_mul_rpow_atTop zero_lt_one _).isBigO + · simp_rw [neg_zero, rpow_zero] + refine isBigO_const_of_tendsto (?_ : Tendsto _ _ (𝓝 (1 : ℂ))) one_ne_zero + rw [(by simp : (1 : ℂ) = Real.exp (-0))] + exact (continuous_ofReal.comp (Real.continuous_exp.comp continuous_neg)).continuousWithinAt + +theorem differentiableAt_GammaAux (s : ℂ) (n : ℕ) (h1 : 1 - s.re < n) (h2 : ∀ m : ℕ, s ≠ -m) : + DifferentiableAt ℂ (GammaAux n) s := by + induction' n with n hn generalizing s + · refine (hasDerivAt_GammaIntegral ?_).differentiableAt + rw [Nat.cast_zero] at h1; linarith + · dsimp only [GammaAux] + specialize hn (s + 1) + have a : 1 - (s + 1).re < ↑n := by + rw [Nat.cast_succ] at h1; rw [Complex.add_re, Complex.one_re]; linarith + have b : ∀ m : ℕ, s + 1 ≠ -m := by + intro m; have := h2 (1 + m) + contrapose! this + rw [← eq_sub_iff_add_eq] at this + simpa using this + refine DifferentiableAt.div (DifferentiableAt.comp _ (hn a b) ?_) ?_ ?_ + · rw [differentiableAt_add_const_iff (1 : ℂ)]; exact differentiableAt_id + · exact differentiableAt_id + · simpa using h2 0 + +theorem differentiableAt_Gamma (s : ℂ) (hs : ∀ m : ℕ, s ≠ -m) : DifferentiableAt ℂ Gamma s := by + let n := ⌊1 - s.re⌋₊ + 1 + have hn : 1 - s.re < n := mod_cast Nat.lt_floor_add_one (1 - s.re) + apply (differentiableAt_GammaAux s n hn hs).congr_of_eventuallyEq + let S := {t : ℂ | 1 - t.re < n} + have : S ∈ 𝓝 s := by + rw [mem_nhds_iff]; use S + refine ⟨Subset.rfl, ?_, hn⟩ + have : S = re ⁻¹' Ioi (1 - n : ℝ) := by + ext; rw [preimage, Ioi, mem_setOf_eq, mem_setOf_eq, mem_setOf_eq]; exact sub_lt_comm + rw [this] + exact Continuous.isOpen_preimage continuous_re _ isOpen_Ioi + apply eventuallyEq_of_mem this + intro t ht; rw [mem_setOf_eq] at ht + apply Gamma_eq_GammaAux; linarith + +end GammaHasDeriv + +/-- At `s = 0`, the Gamma function has a simple pole with residue 1. -/ +theorem tendsto_self_mul_Gamma_nhds_zero : Tendsto (fun z : ℂ => z * Gamma z) (𝓝[≠] 0) (𝓝 1) := by + rw [show 𝓝 (1 : ℂ) = 𝓝 (Gamma (0 + 1)) by simp only [zero_add, Complex.Gamma_one]] + convert (Tendsto.mono_left _ nhdsWithin_le_nhds).congr' + (eventuallyEq_of_mem self_mem_nhdsWithin Complex.Gamma_add_one) using 1 + refine ContinuousAt.comp (g := Gamma) ?_ (continuous_id.add continuous_const).continuousAt + refine (Complex.differentiableAt_Gamma _ fun m => ?_).continuousAt + rw [zero_add, ← ofReal_natCast, ← ofReal_neg, ← ofReal_one, Ne, ofReal_inj] + refine (lt_of_le_of_lt ?_ zero_lt_one).ne' + exact neg_nonpos.mpr (Nat.cast_nonneg _) + +end Complex + +namespace Real + +theorem differentiableAt_Gamma {s : ℝ} (hs : ∀ m : ℕ, s ≠ -m) : DifferentiableAt ℝ Gamma s := by + refine (Complex.differentiableAt_Gamma _ ?_).hasDerivAt.real_of_complex.differentiableAt + simp_rw [← Complex.ofReal_natCast, ← Complex.ofReal_neg, Ne, Complex.ofReal_inj] + exact hs + +end Real From 6bcc518e8f9569a23b8cd02942066bcb5a6aae1b Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Thu, 17 Oct 2024 04:44:25 +0000 Subject: [PATCH 046/505] chore: use funext_iff instead of the alias Function.funext_iff (#17847) Not sure why we were using this alias in the first place. --- Archive/Wiedijk100Theorems/Partition.lean | 2 +- Mathlib/Algebra/BigOperators/Ring.lean | 2 +- Mathlib/Algebra/Divisibility/Prod.lean | 2 +- Mathlib/Algebra/Group/Pi/Lemmas.lean | 4 ++-- Mathlib/Algebra/Homology/Functor.lean | 4 ++-- Mathlib/Algebra/LinearRecurrence.lean | 2 +- Mathlib/Algebra/Module/ZLattice/Basic.lean | 2 +- Mathlib/Algebra/Order/BigOperators/Group/Finset.lean | 4 ++-- Mathlib/Algebra/Polynomial/Smeval.lean | 2 +- Mathlib/Analysis/Calculus/LagrangeMultipliers.lean | 4 ++-- Mathlib/Analysis/InnerProductSpace/OfNorm.lean | 2 +- .../Analysis/Normed/Operator/WeakOperatorTopology.lean | 2 +- Mathlib/Analysis/SpecialFunctions/Complex/Arg.lean | 2 +- Mathlib/CategoryTheory/GradedObject.lean | 2 +- Mathlib/CategoryTheory/Pi/Basic.lean | 2 +- Mathlib/CategoryTheory/Sites/LocallySurjective.lean | 2 +- Mathlib/CategoryTheory/Sites/Types.lean | 2 +- Mathlib/CategoryTheory/Types.lean | 2 +- Mathlib/Combinatorics/Additive/AP/Three/Behrend.lean | 2 +- Mathlib/Combinatorics/HalesJewett.lean | 2 +- Mathlib/Combinatorics/Hall/Basic.lean | 2 +- Mathlib/Combinatorics/SimpleGraph/Basic.lean | 2 +- Mathlib/Combinatorics/SimpleGraph/Circulant.lean | 4 ++-- Mathlib/Combinatorics/SimpleGraph/Turan.lean | 2 +- Mathlib/Computability/TMToPartrec.lean | 2 +- Mathlib/Data/Fin/Basic.lean | 4 ++-- Mathlib/Data/Finset/NAry.lean | 2 +- Mathlib/Data/Finset/Sort.lean | 2 +- Mathlib/Data/Fintype/Basic.lean | 2 +- Mathlib/Data/Fintype/Pi.lean | 6 +++--- Mathlib/Data/FunLike/Basic.lean | 2 +- Mathlib/Data/Matrix/ColumnRowPartitioned.lean | 4 ++-- Mathlib/Data/Multiset/Basic.lean | 2 +- Mathlib/Data/Real/GoldenRatio.lean | 2 +- Mathlib/Data/Set/Image.lean | 2 +- Mathlib/Data/W/Basic.lean | 2 +- Mathlib/FieldTheory/Finite/Polynomial.lean | 2 +- Mathlib/GroupTheory/CommutingProbability.lean | 2 +- Mathlib/GroupTheory/FreeGroup/IsFreeGroup.lean | 2 +- Mathlib/GroupTheory/GroupAction/FixedPoints.lean | 2 +- Mathlib/GroupTheory/HNNExtension.lean | 2 +- Mathlib/GroupTheory/Perm/DomMulAct.lean | 2 +- Mathlib/GroupTheory/PushoutI.lean | 6 +++--- Mathlib/LinearAlgebra/AffineSpace/AffineMap.lean | 2 +- Mathlib/LinearAlgebra/AffineSpace/Independent.lean | 2 +- Mathlib/LinearAlgebra/Dual.lean | 2 +- Mathlib/LinearAlgebra/Matrix/DotProduct.lean | 6 +++--- Mathlib/LinearAlgebra/Pi.lean | 4 ++-- Mathlib/LinearAlgebra/QuadraticForm/Prod.lean | 2 +- Mathlib/Logic/Function/Basic.lean | 2 +- Mathlib/MeasureTheory/Measure/VectorMeasure.lean | 2 +- Mathlib/ModelTheory/Basic.lean | 6 +++--- Mathlib/ModelTheory/ElementaryMaps.lean | 2 +- Mathlib/ModelTheory/Syntax.lean | 2 +- Mathlib/NumberTheory/NumberField/House.lean | 10 +++++----- Mathlib/Order/RelSeries.lean | 2 +- Mathlib/Probability/ProbabilityMassFunction/Basic.lean | 2 +- .../GroupCohomology/LowDegree.lean | 10 +++++----- Mathlib/RingTheory/HahnSeries/Addition.lean | 4 ++-- Mathlib/RingTheory/HahnSeries/Basic.lean | 2 +- Mathlib/RingTheory/HahnSeries/Multiplication.lean | 6 +++--- Mathlib/RingTheory/Ideal/Maps.lean | 2 +- Mathlib/RingTheory/Idempotents.lean | 2 +- Mathlib/RingTheory/WittVector/Defs.lean | 2 +- Mathlib/RingTheory/WittVector/IsPoly.lean | 4 ++-- Mathlib/Topology/Algebra/Module/Basic.lean | 2 +- Mathlib/Topology/Homotopy/Product.lean | 2 +- Mathlib/Topology/LocallyConstant/Basic.lean | 2 +- Mathlib/Topology/TietzeExtension.lean | 2 +- test/instance_diamonds.lean | 8 ++++---- 70 files changed, 101 insertions(+), 101 deletions(-) diff --git a/Archive/Wiedijk100Theorems/Partition.lean b/Archive/Wiedijk100Theorems/Partition.lean index 1045104e58395..689d098bcd228 100644 --- a/Archive/Wiedijk100Theorems/Partition.lean +++ b/Archive/Wiedijk100Theorems/Partition.lean @@ -216,7 +216,7 @@ theorem partialGF_prop (α : Type*) [CommSemiring α] (n : ℕ) (s : Finset ℕ) intro a; exact Nat.lt_irrefl 0 (hs 0 (hp₂.2 0 a)) intro a; exact Nat.lt_irrefl 0 (hs 0 (hp₁.2 0 a)) · rw [← mul_left_inj' hi] - rw [Function.funext_iff] at h + rw [funext_iff] at h exact h.2 i · simp only [φ, mem_filter, mem_finsuppAntidiag, mem_univ, exists_prop, true_and, and_assoc] rintro f ⟨hf, hf₃, hf₄⟩ diff --git a/Mathlib/Algebra/BigOperators/Ring.lean b/Mathlib/Algebra/BigOperators/Ring.lean index 8ffa27de019be..ffe59a97cc2ff 100644 --- a/Mathlib/Algebra/BigOperators/Ring.lean +++ b/Mathlib/Algebra/BigOperators/Ring.lean @@ -168,7 +168,7 @@ theorem prod_add (f g : ι → α) (s : Finset ι) : (fun t _ a _ => a ∈ t) (by simp) (by simp [Classical.em]) - (by simp_rw [mem_filter, Function.funext_iff, eq_iff_iff, mem_pi, mem_insert]; tauto) + (by simp_rw [mem_filter, funext_iff, eq_iff_iff, mem_pi, mem_insert]; tauto) (by simp_rw [Finset.ext_iff, @mem_filter _ _ (id _), mem_powerset]; tauto) (fun a _ ↦ by simp only [prod_ite, filter_attach', prod_map, Function.Embedding.coeFn_mk, diff --git a/Mathlib/Algebra/Divisibility/Prod.lean b/Mathlib/Algebra/Divisibility/Prod.lean index 5eda6b629ea04..355b2cfd6033d 100644 --- a/Mathlib/Algebra/Divisibility/Prod.lean +++ b/Mathlib/Algebra/Divisibility/Prod.lean @@ -33,7 +33,7 @@ instance [DecompositionMonoid G₁] [DecompositionMonoid G₂] : DecompositionMo exact ⟨(a₁, a₂), (a₁', a₂'), ⟨h₁, h₂⟩, ⟨h₁', h₂'⟩, Prod.ext eq₁ eq₂⟩ theorem pi_dvd_iff {x y : ∀ i, G i} : x ∣ y ↔ ∀ i, x i ∣ y i := by - simp_rw [dvd_def, Function.funext_iff, Classical.skolem]; rfl + simp_rw [dvd_def, funext_iff, Classical.skolem]; rfl instance [∀ i, DecompositionMonoid (G i)] : DecompositionMonoid (∀ i, G i) where primal a b c h := by diff --git a/Mathlib/Algebra/Group/Pi/Lemmas.lean b/Mathlib/Algebra/Group/Pi/Lemmas.lean index a344930fb34fa..1dee66ec01d1b 100644 --- a/Mathlib/Algebra/Group/Pi/Lemmas.lean +++ b/Mathlib/Algebra/Group/Pi/Lemmas.lean @@ -70,7 +70,7 @@ def Pi.mulHom {γ : Type w} [∀ i, Mul (f i)] [Mul γ] (g : ∀ i, γ →ₙ* f theorem Pi.mulHom_injective {γ : Type w} [Nonempty I] [∀ i, Mul (f i)] [Mul γ] (g : ∀ i, γ →ₙ* f i) (hg : ∀ i, Function.Injective (g i)) : Function.Injective (Pi.mulHom g) := fun _ _ h => let ⟨i⟩ := ‹Nonempty I› - hg i ((Function.funext_iff.mp h : _) i) + hg i ((funext_iff.mp h : _) i) /-- A family of monoid homomorphisms `f a : γ →* β a` defines a monoid homomorphism `Pi.monoidHom f : γ →* Π a, β a` given by `Pi.monoidHom f x b = f b x`. -/ @@ -337,7 +337,7 @@ theorem SemiconjBy.pi {x y z : ∀ i, f i} (h : ∀ i, SemiconjBy (x i) (y i) (z @[to_additive] theorem Pi.semiconjBy_iff {x y z : ∀ i, f i} : - SemiconjBy x y z ↔ ∀ i, SemiconjBy (x i) (y i) (z i) := Function.funext_iff + SemiconjBy x y z ↔ ∀ i, SemiconjBy (x i) (y i) (z i) := funext_iff @[to_additive] theorem Commute.pi {x y : ∀ i, f i} (h : ∀ i, Commute (x i) (y i)) : Commute x y := .pi h diff --git a/Mathlib/Algebra/Homology/Functor.lean b/Mathlib/Algebra/Homology/Functor.lean index 59cf2538f199c..7560ad47ae5ed 100644 --- a/Mathlib/Algebra/Homology/Functor.lean +++ b/Mathlib/Algebra/Homology/Functor.lean @@ -37,11 +37,11 @@ def asFunctor {T : Type*} [Category T] (C : HomologicalComplex (T ⥤ V) c) : d := fun i j => (C.d i j).app t d_comp_d' := fun i j k _ _ => by have := C.d_comp_d i j k - rw [NatTrans.ext_iff, Function.funext_iff] at this + rw [NatTrans.ext_iff, funext_iff] at this exact this t shape := fun i j h => by have := C.shape _ _ h - rw [NatTrans.ext_iff, Function.funext_iff] at this + rw [NatTrans.ext_iff, funext_iff] at this exact this t } map h := { f := fun i => (C.X i).map h diff --git a/Mathlib/Algebra/LinearRecurrence.lean b/Mathlib/Algebra/LinearRecurrence.lean index b3b5648ee9db0..182a96ce2cf20 100644 --- a/Mathlib/Algebra/LinearRecurrence.lean +++ b/Mathlib/Algebra/LinearRecurrence.lean @@ -138,7 +138,7 @@ def toInit : E.solSpace ≃ₗ[α] Fin E.order → α where simp invFun u := ⟨E.mkSol u, E.is_sol_mkSol u⟩ left_inv u := by ext n; symm; apply E.eq_mk_of_is_sol_of_eq_init u.2; intro k; rfl - right_inv u := Function.funext_iff.mpr fun n ↦ E.mkSol_eq_init u n + right_inv u := funext_iff.mpr fun n ↦ E.mkSol_eq_init u n /-- Two solutions are equal iff they are equal on `range E.order`. -/ theorem sol_eq_of_eq_init (u v : ℕ → α) (hu : E.IsSolution u) (hv : E.IsSolution v) : diff --git a/Mathlib/Algebra/Module/ZLattice/Basic.lean b/Mathlib/Algebra/Module/ZLattice/Basic.lean index a436852b2ab88..f5d850e0a404e 100644 --- a/Mathlib/Algebra/Module/ZLattice/Basic.lean +++ b/Mathlib/Algebra/Module/ZLattice/Basic.lean @@ -276,7 +276,7 @@ theorem discreteTopology_pi_basisFun [Finite ι] : refine discreteTopology_iff_isOpen_singleton_zero.mpr ⟨Metric.ball 0 1, Metric.isOpen_ball, ?_⟩ ext x rw [Set.mem_preimage, mem_ball_zero_iff, pi_norm_lt_iff zero_lt_one, Set.mem_singleton_iff] - simp_rw [← coe_eq_zero, Function.funext_iff, Pi.zero_apply, Real.norm_eq_abs] + simp_rw [← coe_eq_zero, funext_iff, Pi.zero_apply, Real.norm_eq_abs] refine forall_congr' (fun i => ?_) rsuffices ⟨y, hy⟩ : ∃ (y : ℤ), (y : ℝ) = (x : ι → ℝ) i · rw [← hy, ← Int.cast_abs, ← Int.cast_one, Int.cast_lt, Int.abs_lt_one_iff, Int.cast_eq_zero] diff --git a/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean b/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean index d5cdc6b8f86e7..8593406fe5e67 100644 --- a/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean +++ b/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean @@ -519,11 +519,11 @@ lemma one_le_prod (hf : 1 ≤ f) : 1 ≤ ∏ i, f i := Finset.one_le_prod' fun _ @[to_additive] lemma prod_eq_one_iff_of_one_le (hf : 1 ≤ f) : ∏ i, f i = 1 ↔ f = 1 := - (Finset.prod_eq_one_iff_of_one_le' fun i _ ↦ hf i).trans <| by simp [Function.funext_iff] + (Finset.prod_eq_one_iff_of_one_le' fun i _ ↦ hf i).trans <| by simp [funext_iff] @[to_additive] lemma prod_eq_one_iff_of_le_one (hf : f ≤ 1) : ∏ i, f i = 1 ↔ f = 1 := - (Finset.prod_eq_one_iff_of_le_one' fun i _ ↦ hf i).trans <| by simp [Function.funext_iff] + (Finset.prod_eq_one_iff_of_le_one' fun i _ ↦ hf i).trans <| by simp [funext_iff] end OrderedCommMonoid diff --git a/Mathlib/Algebra/Polynomial/Smeval.lean b/Mathlib/Algebra/Polynomial/Smeval.lean index d24a7089bfc93..91d57a909fcc3 100644 --- a/Mathlib/Algebra/Polynomial/Smeval.lean +++ b/Mathlib/Algebra/Polynomial/Smeval.lean @@ -133,7 +133,7 @@ theorem smeval.linearMap_apply : smeval.linearMap R x p = p.smeval x := rfl theorem leval_coe_eq_smeval {R : Type*} [Semiring R] (r : R) : ⇑(leval r) = fun p => p.smeval r := by - rw [Function.funext_iff] + rw [funext_iff] intro rw [leval_apply, smeval_def, eval_eq_sum] rfl diff --git a/Mathlib/Analysis/Calculus/LagrangeMultipliers.lean b/Mathlib/Analysis/Calculus/LagrangeMultipliers.lean index faa6e1df7770b..64e2efb39c875 100644 --- a/Mathlib/Analysis/Calculus/LagrangeMultipliers.lean +++ b/Mathlib/Analysis/Calculus/LagrangeMultipliers.lean @@ -106,7 +106,7 @@ theorem IsLocalExtrOn.exists_multipliers_of_hasStrictFDerivAt {ι : Type*} [Fint ∃ (Λ : ι → ℝ) (Λ₀ : ℝ), (Λ, Λ₀) ≠ 0 ∧ (∑ i, Λ i • f' i) + Λ₀ • φ' = 0 := by letI := Classical.decEq ι replace hextr : IsLocalExtrOn φ {x | (fun i => f i x) = fun i => f i x₀} x₀ := by - simpa only [Function.funext_iff] using hextr + simpa only [funext_iff] using hextr rcases hextr.exists_linear_map_of_hasStrictFDerivAt (hasStrictFDerivAt_pi.2 fun i => hf' i) hφ' with ⟨Λ, Λ₀, h0, hsum⟩ @@ -132,5 +132,5 @@ theorem IsLocalExtrOn.linear_dependent_of_hasStrictFDerivAt {ι : Type*} [Finite rcases hextr.exists_multipliers_of_hasStrictFDerivAt hf' hφ' with ⟨Λ, Λ₀, hΛ, hΛf⟩ refine ⟨Option.elim' Λ₀ Λ, ?_, ?_⟩ · simpa [add_comm] using hΛf - · simpa only [Function.funext_iff, not_and_or, or_comm, Option.exists, Prod.mk_eq_zero, Ne, + · simpa only [funext_iff, not_and_or, or_comm, Option.exists, Prod.mk_eq_zero, Ne, not_forall] using hΛ diff --git a/Mathlib/Analysis/InnerProductSpace/OfNorm.lean b/Mathlib/Analysis/InnerProductSpace/OfNorm.lean index e2e23c0b46a9d..b62522a0e45d6 100644 --- a/Mathlib/Analysis/InnerProductSpace/OfNorm.lean +++ b/Mathlib/Analysis/InnerProductSpace/OfNorm.lean @@ -194,7 +194,7 @@ private theorem rat_prop (r : ℚ) : innerProp' E (r : 𝕜) := by private theorem real_prop (r : ℝ) : innerProp' E (r : 𝕜) := by intro x y revert r - rw [← Function.funext_iff] + rw [← funext_iff] refine Rat.isDenseEmbedding_coe_real.dense.equalizer ?_ ?_ (funext fun X => ?_) · exact (continuous_ofReal.smul continuous_const).inner_ continuous_const · exact (continuous_conj.comp continuous_ofReal).mul continuous_const diff --git a/Mathlib/Analysis/Normed/Operator/WeakOperatorTopology.lean b/Mathlib/Analysis/Normed/Operator/WeakOperatorTopology.lean index 45dfa1892cbcc..202618cf94cd6 100644 --- a/Mathlib/Analysis/Normed/Operator/WeakOperatorTopology.lean +++ b/Mathlib/Analysis/Normed/Operator/WeakOperatorTopology.lean @@ -175,7 +175,7 @@ lemma continuous_of_dual_apply_continuous {α : Type*} [TopologicalSpace α] {g lemma embedding_inducingFn : Embedding (inducingFn 𝕜 E F) := by refine Function.Injective.embedding_induced fun A B hAB => ?_ rw [ContinuousLinearMapWOT.ext_dual_iff] - simpa [Function.funext_iff] using hAB + simpa [funext_iff] using hAB open Filter in /-- The defining property of the weak operator topology: a function `f` tends to diff --git a/Mathlib/Analysis/SpecialFunctions/Complex/Arg.lean b/Mathlib/Analysis/SpecialFunctions/Complex/Arg.lean index b7b5fa601494d..0086e643a19b8 100644 --- a/Mathlib/Analysis/SpecialFunctions/Complex/Arg.lean +++ b/Mathlib/Analysis/SpecialFunctions/Complex/Arg.lean @@ -597,7 +597,7 @@ theorem continuousAt_arg_coe_angle (h : x ≠ 0) : ContinuousAt ((↑) ∘ arg : by_cases hs : x ∈ slitPlane · exact Real.Angle.continuous_coe.continuousAt.comp (continuousAt_arg hs) · rw [← Function.comp_id (((↑) : ℝ → Real.Angle) ∘ arg), - (Function.funext_iff.2 fun _ => (neg_neg _).symm : (id : ℂ → ℂ) = Neg.neg ∘ Neg.neg), ← + (funext_iff.2 fun _ => (neg_neg _).symm : (id : ℂ → ℂ) = Neg.neg ∘ Neg.neg), ← Function.comp_assoc] refine ContinuousAt.comp ?_ continuous_neg.continuousAt suffices ContinuousAt (Function.update (((↑) ∘ arg) ∘ Neg.neg : ℂ → Real.Angle) 0 π) (-x) by diff --git a/Mathlib/CategoryTheory/GradedObject.lean b/Mathlib/CategoryTheory/GradedObject.lean index c3fd1e2c4e0aa..b5344653826c9 100644 --- a/Mathlib/CategoryTheory/GradedObject.lean +++ b/Mathlib/CategoryTheory/GradedObject.lean @@ -158,7 +158,7 @@ abbrev comap {I J : Type*} (h : J → I) : GradedObject I C ⥤ GradedObject J C -- Porting note: added to ease the port, this is a special case of `Functor.eqToHom_proj` @[simp] theorem eqToHom_proj {I : Type*} {x x' : GradedObject I C} (h : x = x') (i : I) : - (eqToHom h : x ⟶ x') i = eqToHom (Function.funext_iff.mp h i) := by + (eqToHom h : x ⟶ x') i = eqToHom (funext_iff.mp h i) := by subst h rfl diff --git a/Mathlib/CategoryTheory/Pi/Basic.lean b/Mathlib/CategoryTheory/Pi/Basic.lean index beeb60fcf3da4..8515247f7d5ba 100644 --- a/Mathlib/CategoryTheory/Pi/Basic.lean +++ b/Mathlib/CategoryTheory/Pi/Basic.lean @@ -206,7 +206,7 @@ section EqToHom @[simp] theorem eqToHom_proj {x x' : ∀ i, C i} (h : x = x') (i : I) : - (eqToHom h : x ⟶ x') i = eqToHom (Function.funext_iff.mp h i) := by + (eqToHom h : x ⟶ x') i = eqToHom (funext_iff.mp h i) := by subst h rfl diff --git a/Mathlib/CategoryTheory/Sites/LocallySurjective.lean b/Mathlib/CategoryTheory/Sites/LocallySurjective.lean index 06a7da7304b54..289d2db4e05a6 100644 --- a/Mathlib/CategoryTheory/Sites/LocallySurjective.lean +++ b/Mathlib/CategoryTheory/Sites/LocallySurjective.lean @@ -93,7 +93,7 @@ instance {F G : Cᵒᵖ ⥤ A} (f : F ⟶ G) [IsLocallySurjective J f] : theorem isLocallySurjective_iff_imagePresheaf_sheafify_eq_top {F G : Cᵒᵖ ⥤ A} (f : F ⟶ G) : IsLocallySurjective J f ↔ (imagePresheaf (whiskerRight f (forget A))).sheafify J = ⊤ := by - simp only [Subpresheaf.ext_iff, Function.funext_iff, Set.ext_iff, top_subpresheaf_obj, + simp only [Subpresheaf.ext_iff, funext_iff, Set.ext_iff, top_subpresheaf_obj, Set.top_eq_univ, Set.mem_univ, iff_true] exact ⟨fun H _ => H.imageSieve_mem, fun H => ⟨H _⟩⟩ diff --git a/Mathlib/CategoryTheory/Sites/Types.lean b/Mathlib/CategoryTheory/Sites/Types.lean index b1ba30f51c8e2..ecbb98c1d2d6f 100644 --- a/Mathlib/CategoryTheory/Sites/Types.lean +++ b/Mathlib/CategoryTheory/Sites/Types.lean @@ -172,6 +172,6 @@ theorem typesGrothendieckTopology_eq_canonical : have : (fun _ => ULift.up true) = fun _ => ULift.up false := (hs PUnit fun _ => x).isSeparatedFor.ext fun β f hf => funext fun y => hsx.elim <| S.2 hf fun _ => y - simp [Function.funext_iff] at this + simp [funext_iff] at this end CategoryTheory diff --git a/Mathlib/CategoryTheory/Types.lean b/Mathlib/CategoryTheory/Types.lean index a084aa296a6af..01484dc45c6dc 100644 --- a/Mathlib/CategoryTheory/Types.lean +++ b/Mathlib/CategoryTheory/Types.lean @@ -117,7 +117,7 @@ lemma sections_property {F : J ⥤ Type w} (s : (F.sections : Type _)) s.property f lemma sections_ext_iff {F : J ⥤ Type w} {x y : F.sections} : x = y ↔ ∀ j, x.val j = y.val j := - Subtype.ext_iff.trans Function.funext_iff + Subtype.ext_iff.trans funext_iff variable (J) diff --git a/Mathlib/Combinatorics/Additive/AP/Three/Behrend.lean b/Mathlib/Combinatorics/Additive/AP/Three/Behrend.lean index 076f7c450bcde..14728c20ba579 100644 --- a/Mathlib/Combinatorics/Additive/AP/Three/Behrend.lean +++ b/Mathlib/Combinatorics/Additive/AP/Three/Behrend.lean @@ -102,7 +102,7 @@ quadrant. -/ def sphere (n d k : ℕ) : Finset (Fin n → ℕ) := (box n d).filter fun x => ∑ i, x i ^ 2 = k -theorem sphere_zero_subset : sphere n d 0 ⊆ 0 := fun x => by simp [sphere, Function.funext_iff] +theorem sphere_zero_subset : sphere n d 0 ⊆ 0 := fun x => by simp [sphere, funext_iff] @[simp] theorem sphere_zero_right (n k : ℕ) : sphere (n + 1) 0 k = ∅ := by simp [sphere] diff --git a/Mathlib/Combinatorics/HalesJewett.lean b/Mathlib/Combinatorics/HalesJewett.lean index 04d6ee1658e63..76bb5d7812114 100644 --- a/Mathlib/Combinatorics/HalesJewett.lean +++ b/Mathlib/Combinatorics/HalesJewett.lean @@ -138,7 +138,7 @@ variable {η' α' ι' : Type*} def reindex (l : Subspace η α ι) (eη : η ≃ η') (eα : α ≃ α') (eι : ι ≃ ι') : Subspace η' α' ι' where idxFun i := (l.idxFun <| eι.symm i).map eα eη proper e := (eι.exists_congr fun i ↦ by cases h : idxFun l i <;> - simp [*, Function.funext_iff, Equiv.eq_symm_apply]).1 <| l.proper <| eη.symm e + simp [*, funext_iff, Equiv.eq_symm_apply]).1 <| l.proper <| eη.symm e @[simp] lemma reindex_apply (l : Subspace η α ι) (eη : η ≃ η') (eα : α ≃ α') (eι : ι ≃ ι') (x i) : l.reindex eη eα eι x i = eα (l (eα.symm ∘ x ∘ eη) <| eι.symm i) := by diff --git a/Mathlib/Combinatorics/Hall/Basic.lean b/Mathlib/Combinatorics/Hall/Basic.lean index 70a89a8ee9790..0d0b9a82657a9 100644 --- a/Mathlib/Combinatorics/Hall/Basic.lean +++ b/Mathlib/Combinatorics/Hall/Basic.lean @@ -100,7 +100,7 @@ instance hallMatchingsOn.finite {ι : Type u} {α : Type v} (t : ι → Finset apply Finite.of_injective g intro f f' h ext a - rw [Function.funext_iff] at h + rw [funext_iff] at h simpa [g] using h a /-- This is the version of **Hall's Marriage Theorem** in terms of indexed diff --git a/Mathlib/Combinatorics/SimpleGraph/Basic.lean b/Mathlib/Combinatorics/SimpleGraph/Basic.lean index 008ba18f37ef6..69de374b0e114 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Basic.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Basic.lean @@ -355,7 +355,7 @@ instance [Subsingleton V] : Unique (SimpleGraph V) where uniq G := by ext a b; have := Subsingleton.elim a b; simp [this] instance [Nontrivial V] : Nontrivial (SimpleGraph V) := - ⟨⟨⊥, ⊤, fun h ↦ not_subsingleton V ⟨by simpa only [← adj_inj, Function.funext_iff, bot_adj, + ⟨⟨⊥, ⊤, fun h ↦ not_subsingleton V ⟨by simpa only [← adj_inj, funext_iff, bot_adj, top_adj, ne_eq, eq_iff_iff, false_iff, not_not] using h⟩⟩⟩ section Decidable diff --git a/Mathlib/Combinatorics/SimpleGraph/Circulant.lean b/Mathlib/Combinatorics/SimpleGraph/Circulant.lean index 121f720ec7a12..99106b3256d7b 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Circulant.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Circulant.lean @@ -72,11 +72,11 @@ theorem cycleGraph_zero_eq_top : cycleGraph 0 = ⊤ := Subsingleton.elim _ _ theorem cycleGraph_one_eq_top : cycleGraph 1 = ⊤ := Subsingleton.elim _ _ theorem cycleGraph_two_eq_top : cycleGraph 2 = ⊤ := by - simp only [SimpleGraph.ext_iff, Function.funext_iff] + simp only [SimpleGraph.ext_iff, funext_iff] decide theorem cycleGraph_three_eq_top : cycleGraph 3 = ⊤ := by - simp only [SimpleGraph.ext_iff, Function.funext_iff] + simp only [SimpleGraph.ext_iff, funext_iff] decide theorem cycleGraph_one_adj {u v : Fin 1} : ¬(cycleGraph 1).Adj u v := by diff --git a/Mathlib/Combinatorics/SimpleGraph/Turan.lean b/Mathlib/Combinatorics/SimpleGraph/Turan.lean index 45760dc0bd670..08ab8e530a2c0 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Turan.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Turan.lean @@ -72,7 +72,7 @@ lemma turanGraph_zero : turanGraph n 0 = ⊤ := by @[simp] theorem turanGraph_eq_top : turanGraph n r = ⊤ ↔ r = 0 ∨ n ≤ r := by - simp_rw [SimpleGraph.ext_iff, Function.funext_iff, turanGraph, top_adj, eq_iff_iff, not_iff_not] + simp_rw [SimpleGraph.ext_iff, funext_iff, turanGraph, top_adj, eq_iff_iff, not_iff_not] refine ⟨fun h ↦ ?_, ?_⟩ · contrapose! h use ⟨0, (Nat.pos_of_ne_zero h.1).trans h.2⟩, ⟨r, h.2⟩ diff --git a/Mathlib/Computability/TMToPartrec.lean b/Mathlib/Computability/TMToPartrec.lean index a19ae0a5453fd..440884be0810c 100644 --- a/Mathlib/Computability/TMToPartrec.lean +++ b/Mathlib/Computability/TMToPartrec.lean @@ -883,7 +883,7 @@ instance Λ'.instInhabited : Inhabited Λ' := instance Λ'.instDecidableEq : DecidableEq Λ' := fun a b => by induction a generalizing b <;> cases b <;> first | apply Decidable.isFalse; rintro ⟨⟨⟩⟩; done - | exact decidable_of_iff' _ (by simp [Function.funext_iff]; rfl) + | exact decidable_of_iff' _ (by simp [funext_iff]; rfl) /-- The type of TM2 statements used by this machine. -/ def Stmt' := diff --git a/Mathlib/Data/Fin/Basic.lean b/Mathlib/Data/Fin/Basic.lean index 7b7b730d2913b..715342aa56963 100644 --- a/Mathlib/Data/Fin/Basic.lean +++ b/Mathlib/Data/Fin/Basic.lean @@ -171,7 +171,7 @@ then they coincide (in the heq sense). -/ protected theorem heq_fun_iff {α : Sort*} {k l : ℕ} (h : k = l) {f : Fin k → α} {g : Fin l → α} : HEq f g ↔ ∀ i : Fin k, f i = g ⟨(i : ℕ), h ▸ i.2⟩ := by subst h - simp [Function.funext_iff] + simp [funext_iff] /-- Assume `k = l` and `k' = l'`. If two functions `Fin k → Fin k' → α` and `Fin l → Fin l' → α` are equal on each pair, @@ -181,7 +181,7 @@ protected theorem heq_fun₂_iff {α : Sort*} {k l k' l' : ℕ} (h : k = l) (h' HEq f g ↔ ∀ (i : Fin k) (j : Fin k'), f i j = g ⟨(i : ℕ), h ▸ i.2⟩ ⟨(j : ℕ), h' ▸ j.2⟩ := by subst h subst h' - simp [Function.funext_iff] + simp [funext_iff] /-- Two elements of `Fin k` and `Fin l` are heq iff their values in `ℕ` coincide. This requires `k = l`. For the left implication without this assumption, see `val_eq_val_of_heq`. -/ diff --git a/Mathlib/Data/Finset/NAry.lean b/Mathlib/Data/Finset/NAry.lean index 97c9c4284cbd7..db73633e28a15 100644 --- a/Mathlib/Data/Finset/NAry.lean +++ b/Mathlib/Data/Finset/NAry.lean @@ -582,7 +582,7 @@ variable {ι : Type*} {α β γ : ι → Type*} [DecidableEq ι] [Fintype ι] [ lemma piFinset_image₂ (f : ∀ i, α i → β i → γ i) (s : ∀ i, Finset (α i)) (t : ∀ i, Finset (β i)) : piFinset (fun i ↦ image₂ (f i) (s i) (t i)) = image₂ (fun a b i ↦ f _ (a i) (b i)) (piFinset s) (piFinset t) := by - ext; simp only [mem_piFinset, mem_image₂, Classical.skolem, forall_and, Function.funext_iff] + ext; simp only [mem_piFinset, mem_image₂, Classical.skolem, forall_and, funext_iff] end Fintype diff --git a/Mathlib/Data/Finset/Sort.lean b/Mathlib/Data/Finset/Sort.lean index 361e4f3721a7f..923e35810f214 100644 --- a/Mathlib/Data/Finset/Sort.lean +++ b/Mathlib/Data/Finset/Sort.lean @@ -219,7 +219,7 @@ theorem orderEmbOfFin_unique {s : Finset α} {k : ℕ} (h : s.card = k) {f : Fin the increasing bijection `orderEmbOfFin s h`. -/ theorem orderEmbOfFin_unique' {s : Finset α} {k : ℕ} (h : s.card = k) {f : Fin k ↪o α} (hfs : ∀ x, f x ∈ s) : f = s.orderEmbOfFin h := - RelEmbedding.ext <| Function.funext_iff.1 <| orderEmbOfFin_unique h hfs f.strictMono + RelEmbedding.ext <| funext_iff.1 <| orderEmbOfFin_unique h hfs f.strictMono /-- Two parametrizations `orderEmbOfFin` of the same set take the same value on `i` and `j` if and only if `i = j`. Since they can be defined on a priori not defeq types `Fin k` and `Fin l` diff --git a/Mathlib/Data/Fintype/Basic.lean b/Mathlib/Data/Fintype/Basic.lean index 0ff3b973f0141..997a6dd16aeb4 100644 --- a/Mathlib/Data/Fintype/Basic.lean +++ b/Mathlib/Data/Fintype/Basic.lean @@ -376,7 +376,7 @@ namespace Fintype instance decidablePiFintype {α} {β : α → Type*} [∀ a, DecidableEq (β a)] [Fintype α] : DecidableEq (∀ a, β a) := fun f g => decidable_of_iff (∀ a ∈ @Fintype.elems α _, f a = g a) - (by simp [Function.funext_iff, Fintype.complete]) + (by simp [funext_iff, Fintype.complete]) instance decidableForallFintype {p : α → Prop} [DecidablePred p] [Fintype α] : Decidable (∀ a, p a) := diff --git a/Mathlib/Data/Fintype/Pi.lean b/Mathlib/Data/Fintype/Pi.lean index 99a88bda7004a..c32f9bbfd62b6 100644 --- a/Mathlib/Data/Fintype/Pi.lean +++ b/Mathlib/Data/Fintype/Pi.lean @@ -25,7 +25,7 @@ analogue of `Finset.pi` where the base finset is `univ` (but formally they are n there is an additional condition `i ∈ Finset.univ` in the `Finset.pi` definition). -/ def piFinset (t : ∀ a, Finset (δ a)) : Finset (∀ a, δ a) := (Finset.univ.pi t).map ⟨fun f a => f a (mem_univ a), fun _ _ => - by simp (config := {contextual := true}) [Function.funext_iff]⟩ + by simp (config := {contextual := true}) [funext_iff]⟩ @[simp] theorem mem_piFinset {t : ∀ a, Finset (δ a)} {f : ∀ a, δ a} : f ∈ piFinset t ↔ ∀ a, f a ∈ t a := by @@ -69,7 +69,7 @@ lemma piFinset_of_isEmpty [IsEmpty α] (s : ∀ a, Finset (γ a)) : piFinset s = @[simp] theorem piFinset_singleton (f : ∀ i, δ i) : piFinset (fun i => {f i} : ∀ i, Finset (δ i)) = {f} := - ext fun _ => by simp only [Function.funext_iff, Fintype.mem_piFinset, mem_singleton] + ext fun _ => by simp only [funext_iff, Fintype.mem_piFinset, mem_singleton] theorem piFinset_subsingleton {f : ∀ i, Finset (δ i)} (hf : ∀ i, (f i : Set (δ i)).Subsingleton) : (Fintype.piFinset f : Set (∀ i, δ i)).Subsingleton := fun _ ha _ hb => @@ -83,7 +83,7 @@ theorem piFinset_disjoint_of_disjoint (t₁ t₂ : ∀ a, Finset (δ a)) {a : α lemma piFinset_image [∀ a, DecidableEq (δ a)] (f : ∀ a, γ a → δ a) (s : ∀ a, Finset (γ a)) : piFinset (fun a ↦ (s a).image (f a)) = (piFinset s).image fun b a ↦ f _ (b a) := by - ext; simp only [mem_piFinset, mem_image, Classical.skolem, forall_and, Function.funext_iff] + ext; simp only [mem_piFinset, mem_image, Classical.skolem, forall_and, funext_iff] lemma eval_image_piFinset_subset (t : ∀ a, Finset (δ a)) (a : α) [DecidableEq (δ a)] : ((piFinset t).image fun f ↦ f a) ⊆ t a := image_subset_iff.2 fun _x hx ↦ mem_piFinset.1 hx _ diff --git a/Mathlib/Data/FunLike/Basic.lean b/Mathlib/Data/FunLike/Basic.lean index 3a04435a85ca6..afcc1b7e13a25 100644 --- a/Mathlib/Data/FunLike/Basic.lean +++ b/Mathlib/Data/FunLike/Basic.lean @@ -194,7 +194,7 @@ theorem ext (f g : F) (h : ∀ x : α, f x = g x) : f = g := DFunLike.coe_injective' (funext h) theorem ext_iff {f g : F} : f = g ↔ ∀ x, f x = g x := - coe_fn_eq.symm.trans Function.funext_iff + coe_fn_eq.symm.trans funext_iff protected theorem congr_fun {f g : F} (h₁ : f = g) (x : α) : f x = g x := congr_fun (congr_arg _ h₁) x diff --git a/Mathlib/Data/Matrix/ColumnRowPartitioned.lean b/Mathlib/Data/Matrix/ColumnRowPartitioned.lean index 1fef44043ba64..67906c2b869f0 100644 --- a/Mathlib/Data/Matrix/ColumnRowPartitioned.lean +++ b/Mathlib/Data/Matrix/ColumnRowPartitioned.lean @@ -108,12 +108,12 @@ lemma fromRows_toRows (A : Matrix (m₁ ⊕ m₂) n R) : fromRows A.toRows₁ A. lemma fromRows_inj : Function.Injective2 (@fromRows R m₁ m₂ n) := by intros x1 x2 y1 y2 - simp only [Function.funext_iff, ← Matrix.ext_iff] + simp only [funext_iff, ← Matrix.ext_iff] aesop lemma fromColumns_inj : Function.Injective2 (@fromColumns R m n₁ n₂) := by intros x1 x2 y1 y2 - simp only [Function.funext_iff, ← Matrix.ext_iff] + simp only [funext_iff, ← Matrix.ext_iff] aesop lemma fromColumns_ext_iff (A₁ : Matrix m n₁ R) (A₂ : Matrix m n₂ R) (B₁ : Matrix m n₁ R) diff --git a/Mathlib/Data/Multiset/Basic.lean b/Mathlib/Data/Multiset/Basic.lean index 7dc4316f0e993..8f15a338625d1 100644 --- a/Mathlib/Data/Multiset/Basic.lean +++ b/Mathlib/Data/Multiset/Basic.lean @@ -1421,7 +1421,7 @@ instance decidableDforallMultiset {p : ∀ a ∈ m, Prop} [_hp : ∀ (a) (h : a /-- decidable equality for functions whose domain is bounded by multisets -/ instance decidableEqPiMultiset {β : α → Type*} [∀ a, DecidableEq (β a)] : DecidableEq (∀ a ∈ m, β a) := fun f g => - decidable_of_iff (∀ (a) (h : a ∈ m), f a h = g a h) (by simp [Function.funext_iff]) + decidable_of_iff (∀ (a) (h : a ∈ m), f a h = g a h) (by simp [funext_iff]) /-- If `p` is a decidable predicate, so is the existence of an element in a multiset satisfying `p`. -/ diff --git a/Mathlib/Data/Real/GoldenRatio.lean b/Mathlib/Data/Real/GoldenRatio.lean index 42e9c2559ca88..2518fadf3188a 100644 --- a/Mathlib/Data/Real/GoldenRatio.lean +++ b/Mathlib/Data/Real/GoldenRatio.lean @@ -201,7 +201,7 @@ theorem Real.coe_fib_eq' : /-- Binet's formula as a dependent equality. -/ theorem Real.coe_fib_eq : ∀ n, (Nat.fib n : ℝ) = (φ ^ n - ψ ^ n) / √5 := by - rw [← Function.funext_iff, Real.coe_fib_eq'] + rw [← funext_iff, Real.coe_fib_eq'] /-- Relationship between the Fibonacci Sequence, Golden Ratio and its conjugate's exponents --/ theorem fib_golden_conj_exp (n : ℕ) : Nat.fib (n + 1) - φ * Nat.fib n = ψ ^ n := by diff --git a/Mathlib/Data/Set/Image.lean b/Mathlib/Data/Set/Image.lean index b0d95b24f9381..071d9c15c3f6c 100644 --- a/Mathlib/Data/Set/Image.lean +++ b/Mathlib/Data/Set/Image.lean @@ -634,7 +634,7 @@ theorem range_subset_iff : range f ⊆ s ↔ ∀ y, f y ∈ s := theorem range_subset_range_iff_exists_comp {f : α → γ} {g : β → γ} : range f ⊆ range g ↔ ∃ h : α → β, f = g ∘ h := by - simp only [range_subset_iff, mem_range, Classical.skolem, Function.funext_iff, (· ∘ ·), eq_comm] + simp only [range_subset_iff, mem_range, Classical.skolem, funext_iff, (· ∘ ·), eq_comm] theorem range_eq_iff (f : α → β) (s : Set β) : range f = s ↔ (∀ a, f a ∈ s) ∧ ∀ b ∈ s, ∃ a, f a = b := by diff --git a/Mathlib/Data/W/Basic.lean b/Mathlib/Data/W/Basic.lean index a2b73e07e8871..7c1f339cc49d0 100644 --- a/Mathlib/Data/W/Basic.lean +++ b/Mathlib/Data/W/Basic.lean @@ -103,7 +103,7 @@ theorem infinite_of_nonempty_of_isEmpty (a b : α) [ha : Nonempty (β a)] [he : · cases' m with m · simp_all · refine congr_arg Nat.succ (ih ?_) - simp_all [Function.funext_iff]⟩ + simp_all [funext_iff]⟩ variable [∀ a : α, Fintype (β a)] diff --git a/Mathlib/FieldTheory/Finite/Polynomial.lean b/Mathlib/FieldTheory/Finite/Polynomial.lean index a5a76e1726feb..36dd816621f74 100644 --- a/Mathlib/FieldTheory/Finite/Polynomial.lean +++ b/Mathlib/FieldTheory/Finite/Polynomial.lean @@ -99,7 +99,7 @@ end CommRing variable [Field K] theorem eval_indicator_apply_eq_zero (a b : σ → K) (h : a ≠ b) : eval a (indicator b) = 0 := by - obtain ⟨i, hi⟩ : ∃ i, a i ≠ b i := by rwa [Ne, Function.funext_iff, not_forall] at h + obtain ⟨i, hi⟩ : ∃ i, a i ≠ b i := by rwa [Ne, funext_iff, not_forall] at h simp only [indicator, map_prod, map_sub, map_one, map_pow, eval_X, eval_C, sub_self, Finset.prod_eq_zero_iff] refine ⟨i, Finset.mem_univ _, ?_⟩ diff --git a/Mathlib/GroupTheory/CommutingProbability.lean b/Mathlib/GroupTheory/CommutingProbability.lean index 3f168c3c28305..df8f7702b7fdd 100644 --- a/Mathlib/GroupTheory/CommutingProbability.lean +++ b/Mathlib/GroupTheory/CommutingProbability.lean @@ -47,7 +47,7 @@ theorem commProb_prod (M' : Type*) [Mul M'] : commProb (M × M') = commProb M * theorem commProb_pi {α : Type*} (i : α → Type*) [Fintype α] [∀ a, Mul (i a)] : commProb (∀ a, i a) = ∏ a, commProb (i a) := by simp_rw [commProb_def, Finset.prod_div_distrib, Finset.prod_pow, ← Nat.cast_prod, - ← Nat.card_pi, Commute, SemiconjBy, Function.funext_iff] + ← Nat.card_pi, Commute, SemiconjBy, funext_iff] congr 2 exact Nat.card_congr ⟨fun x a => ⟨⟨x.1.1 a, x.1.2 a⟩, x.2 a⟩, fun x => ⟨⟨fun a => (x a).1.1, fun a => (x a).1.2⟩, fun a => (x a).2⟩, fun x => rfl, fun x => rfl⟩ diff --git a/Mathlib/GroupTheory/FreeGroup/IsFreeGroup.lean b/Mathlib/GroupTheory/FreeGroup/IsFreeGroup.lean index 5999f4a20b721..e3a6f25954b2b 100644 --- a/Mathlib/GroupTheory/FreeGroup/IsFreeGroup.lean +++ b/Mathlib/GroupTheory/FreeGroup/IsFreeGroup.lean @@ -214,7 +214,7 @@ group extends in a unique way to a homomorphism from `G`. Note that since `IsFreeGroup.lift` is expressed as a bijection, it already expresses the universal property. -/ theorem unique_lift (f : Generators G → H) : ∃! F : G →* H, ∀ a, F (of a) = f a := by - simpa only [Function.funext_iff] using lift.symm.bijective.existsUnique f + simpa only [funext_iff] using lift.symm.bijective.existsUnique f /-- If a group satisfies the universal property of a free group with respect to a given type, then it is free. Here, the universal property is expressed as in `IsFreeGroup.lift` and its diff --git a/Mathlib/GroupTheory/GroupAction/FixedPoints.lean b/Mathlib/GroupTheory/GroupAction/FixedPoints.lean index 070b3e8dd30e2..a8a5bc7b71b55 100644 --- a/Mathlib/GroupTheory/GroupAction/FixedPoints.lean +++ b/Mathlib/GroupTheory/GroupAction/FixedPoints.lean @@ -237,7 +237,7 @@ then `fixedBy α m = Set.univ` implies that `m = 1`. -/ then `fixedBy α m = Set.univ` implies that `m = 1`."] theorem fixedBy_eq_univ_iff_eq_one {m : M} : fixedBy α m = Set.univ ↔ m = 1 := by rw [← (smul_left_injective' (M := M) (α := α)).eq_iff, Set.eq_univ_iff_forall] - simp_rw [Function.funext_iff, one_smul, mem_fixedBy] + simp_rw [funext_iff, one_smul, mem_fixedBy] /-- If the image of the `(fixedBy α g)ᶜ` set by the pointwise action of `h: G` diff --git a/Mathlib/GroupTheory/HNNExtension.lean b/Mathlib/GroupTheory/HNNExtension.lean index 551403b5170c4..a4d0b7d277511 100644 --- a/Mathlib/GroupTheory/HNNExtension.lean +++ b/Mathlib/GroupTheory/HNNExtension.lean @@ -595,7 +595,7 @@ theorem of_injective : Function.Injective (of : G → HNNExtension G A B φ) := (f := ((· • ·) : HNNExtension G A B φ → NormalWord d → NormalWord d)) ?_ intros _ _ h exact eq_of_smul_eq_smul (fun w : NormalWord d => - by simp_all [Function.funext_iff, of_smul_eq_smul]) + by simp_all [funext_iff, of_smul_eq_smul]) namespace ReducedWord diff --git a/Mathlib/GroupTheory/Perm/DomMulAct.lean b/Mathlib/GroupTheory/Perm/DomMulAct.lean index 5d730372bcddf..4db961f223713 100644 --- a/Mathlib/GroupTheory/Perm/DomMulAct.lean +++ b/Mathlib/GroupTheory/Perm/DomMulAct.lean @@ -139,7 +139,7 @@ theorem stabilizer_card': rw [refl_apply, ← Subtype.coe_inj] simp only [φ, Set.val_codRestrict_apply] · intro g - simp only [Function.funext_iff] + simp only [funext_iff] apply forall_congr' intro a simp only [Function.comp_apply, φ, ← Subtype.coe_inj, Set.val_codRestrict_apply] diff --git a/Mathlib/GroupTheory/PushoutI.lean b/Mathlib/GroupTheory/PushoutI.lean index 47fa6da660134..0f672f3da1e25 100644 --- a/Mathlib/GroupTheory/PushoutI.lean +++ b/Mathlib/GroupTheory/PushoutI.lean @@ -152,7 +152,7 @@ def homEquiv : invFun := fun f => lift f.1.1 f.1.2 f.2, left_inv := fun _ => hom_ext (by simp [DFunLike.ext_iff]) (by simp [DFunLike.ext_iff]) - right_inv := fun ⟨⟨_, _⟩, _⟩ => by simp [DFunLike.ext_iff, Function.funext_iff] } + right_inv := fun ⟨⟨_, _⟩, _⟩ => by simp [DFunLike.ext_iff, funext_iff] } /-- The map from the coproduct into the pushout -/ def ofCoprodI : CoprodI G →* PushoutI φ := @@ -596,7 +596,7 @@ theorem of_injective (hφ : ∀ i, Function.Injective (φ i)) (i : ι) : (f := ((· • ·) : PushoutI φ → NormalWord d → NormalWord d)) ?_ intros _ _ h exact eq_of_smul_eq_smul (fun w : NormalWord d => - by simp_all [Function.funext_iff, of_smul_eq_smul]) + by simp_all [funext_iff, of_smul_eq_smul]) theorem base_injective (hφ : ∀ i, Function.Injective (φ i)) : Function.Injective (base φ) := by @@ -607,7 +607,7 @@ theorem base_injective (hφ : ∀ i, Function.Injective (φ i)) : (f := ((· • ·) : PushoutI φ → NormalWord d → NormalWord d)) ?_ intros _ _ h exact eq_of_smul_eq_smul (fun w : NormalWord d => - by simp_all [Function.funext_iff, base_smul_eq_smul]) + by simp_all [funext_iff, base_smul_eq_smul]) section Reduced diff --git a/Mathlib/LinearAlgebra/AffineSpace/AffineMap.lean b/Mathlib/LinearAlgebra/AffineSpace/AffineMap.lean index ef82d9e40a82c..edeab9e15af30 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/AffineMap.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/AffineMap.lean @@ -703,7 +703,7 @@ theorem pi_comp (g : P3 →ᵃ[k] P1) : (pi fp).comp g = pi (fun i => (fp i).com rfl theorem pi_eq_zero : pi fv = 0 ↔ ∀ i, fv i = 0 := by - simp only [AffineMap.ext_iff, Function.funext_iff, pi_apply] + simp only [AffineMap.ext_iff, funext_iff, pi_apply] exact forall_comm theorem pi_zero : pi (fun _ ↦ 0 : (i : ι) → P1 →ᵃ[k] φv i) = 0 := by diff --git a/Mathlib/LinearAlgebra/AffineSpace/Independent.lean b/Mathlib/LinearAlgebra/AffineSpace/Independent.lean index 24d2fdd76c35b..0fe67bc39f3f9 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/Independent.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/Independent.lean @@ -908,7 +908,7 @@ theorem centroid_eq_iff [CharZero k] {n : ℕ} (s : Simplex k P n) {fs₁ fs₂ (affineIndependent_iff_indicator_eq_of_affineCombination_eq k s.points).1 s.independent _ _ _ _ (fs₁.sum_centroidWeightsIndicator_eq_one_of_card_eq_add_one k h₁) (fs₂.sum_centroidWeightsIndicator_eq_one_of_card_eq_add_one k h₂) h - simp_rw [Finset.coe_univ, Set.indicator_univ, Function.funext_iff, + simp_rw [Finset.coe_univ, Set.indicator_univ, funext_iff, Finset.centroidWeightsIndicator_def, Finset.centroidWeights, h₁, h₂] at ha ext i specialize ha i diff --git a/Mathlib/LinearAlgebra/Dual.lean b/Mathlib/LinearAlgebra/Dual.lean index 5d556906e958a..0dfa5a4d643f0 100644 --- a/Mathlib/LinearAlgebra/Dual.lean +++ b/Mathlib/LinearAlgebra/Dual.lean @@ -1228,7 +1228,7 @@ instance (W : Submodule R M) : FunLike (W.dualAnnihilator) M R where coe φ := φ.val coe_injective' φ ψ h := by ext - simp only [Function.funext_iff] at h + simp only [funext_iff] at h exact h _ @[simp] diff --git a/Mathlib/LinearAlgebra/Matrix/DotProduct.lean b/Mathlib/LinearAlgebra/Matrix/DotProduct.lean index 3a313af6368d5..bb06aeea3eee8 100644 --- a/Mathlib/LinearAlgebra/Matrix/DotProduct.lean +++ b/Mathlib/LinearAlgebra/Matrix/DotProduct.lean @@ -88,7 +88,7 @@ variable [Fintype m] [Fintype n] [Fintype p] @[simp] theorem dotProduct_self_eq_zero [LinearOrderedRing R] {v : n → R} : dotProduct v v = 0 ↔ v = 0 := (Finset.sum_eq_zero_iff_of_nonneg fun i _ => mul_self_nonneg (v i)).trans <| by - simp [Function.funext_iff] + simp [funext_iff] section StarOrderedRing @@ -110,13 +110,13 @@ variable [NoZeroDivisors R] @[simp] theorem dotProduct_star_self_eq_zero {v : n → R} : dotProduct (star v) v = 0 ↔ v = 0 := (Fintype.sum_eq_zero_iff_of_nonneg fun _ => star_mul_self_nonneg _).trans <| - by simp [Function.funext_iff, mul_eq_zero] + by simp [funext_iff, mul_eq_zero] /-- Note that this applies to `ℂ` via `RCLike.toStarOrderedRing`. -/ @[simp] theorem dotProduct_self_star_eq_zero {v : n → R} : dotProduct v (star v) = 0 ↔ v = 0 := (Fintype.sum_eq_zero_iff_of_nonneg fun _ => mul_star_self_nonneg _).trans <| - by simp [Function.funext_iff, mul_eq_zero] + by simp [funext_iff, mul_eq_zero] @[simp] lemma conjTranspose_mul_self_eq_zero {n} {A : Matrix m n R} : Aᴴ * A = 0 ↔ A = 0 := diff --git a/Mathlib/LinearAlgebra/Pi.lean b/Mathlib/LinearAlgebra/Pi.lean index 89740b2d0a89e..079b058aab3d1 100644 --- a/Mathlib/LinearAlgebra/Pi.lean +++ b/Mathlib/LinearAlgebra/Pi.lean @@ -462,8 +462,8 @@ def piOptionEquivProd {ι : Type*} {M : Option ι → Type*} [(i : Option ι) [(i : Option ι) → Module R (M i)] : ((i : Option ι) → M i) ≃ₗ[R] M none × ((i : ι) → M (some i)) := { Equiv.piOptionEquivProd with - map_add' := by simp [Function.funext_iff] - map_smul' := by simp [Function.funext_iff] } + map_add' := by simp [funext_iff] + map_smul' := by simp [funext_iff] } variable (ι M) (S : Type*) [Fintype ι] [DecidableEq ι] [Semiring S] [AddCommMonoid M] [Module R M] [Module S M] [SMulCommClass R S M] diff --git a/Mathlib/LinearAlgebra/QuadraticForm/Prod.lean b/Mathlib/LinearAlgebra/QuadraticForm/Prod.lean index 891097a9ae0b6..20c4468b0c051 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/Prod.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/Prod.lean @@ -303,7 +303,7 @@ theorem Equivalent.pi [Fintype ι] {Q : ∀ i, QuadraticMap R (Mᵢ i) P} /-- If a family is anisotropic then its components must be. The converse is not true. -/ theorem anisotropic_of_pi [Fintype ι] {Q : ∀ i, QuadraticMap R (Mᵢ i) P} (h : (pi Q).Anisotropic) : ∀ i, (Q i).Anisotropic := by - simp_rw [Anisotropic, pi_apply, Function.funext_iff, Pi.zero_apply] at h + simp_rw [Anisotropic, pi_apply, funext_iff, Pi.zero_apply] at h intro i x hx classical have := h (Pi.single i x) ?_ i diff --git a/Mathlib/Logic/Function/Basic.lean b/Mathlib/Logic/Function/Basic.lean index 131ecb1c0731d..a7cb654f6a4e7 100644 --- a/Mathlib/Logic/Function/Basic.lean +++ b/Mathlib/Logic/Function/Basic.lean @@ -893,7 +893,7 @@ lemma forall_existsUnique_iff {r : α → β → Prop} : if and only if it is `(f · = ·)` for some function `f`. -/ lemma forall_existsUnique_iff' {r : α → β → Prop} : (∀ a, ∃! b, r a b) ↔ ∃ f : α → β, r = (f · = ·) := by - simp [forall_existsUnique_iff, Function.funext_iff] + simp [forall_existsUnique_iff, funext_iff] /-- A symmetric relation `r : α → α → Prop` is "function-like" (for each `a` there exists a unique `b` such that `r a b`) diff --git a/Mathlib/MeasureTheory/Measure/VectorMeasure.lean b/Mathlib/MeasureTheory/Measure/VectorMeasure.lean index d6bc3aa0248e0..06d029fc3a63b 100644 --- a/Mathlib/MeasureTheory/Measure/VectorMeasure.lean +++ b/Mathlib/MeasureTheory/Measure/VectorMeasure.lean @@ -96,7 +96,7 @@ theorem coe_injective : @Function.Injective (VectorMeasure α M) (Set α → M) congr theorem ext_iff' (v w : VectorMeasure α M) : v = w ↔ ∀ i : Set α, v i = w i := by - rw [← coe_injective.eq_iff, Function.funext_iff] + rw [← coe_injective.eq_iff, funext_iff] theorem ext_iff (v w : VectorMeasure α M) : v = w ↔ ∀ i : Set α, MeasurableSet i → v i = w i := by constructor diff --git a/Mathlib/ModelTheory/Basic.lean b/Mathlib/ModelTheory/Basic.lean index baa0ada8d4271..6846aee6eb65c 100644 --- a/Mathlib/ModelTheory/Basic.lean +++ b/Mathlib/ModelTheory/Basic.lean @@ -379,7 +379,7 @@ instance funLike : FunLike (M ↪[L] N) M N where cases g congr ext x - exact Function.funext_iff.1 h x + exact funext_iff.1 h x instance embeddingLike : EmbeddingLike (M ↪[L] N) M N where injective' f := f.toEmbedding.injective @@ -416,7 +416,7 @@ theorem coe_injective : @Function.Injective (M ↪[L] N) (M → N) (↑) cases g congr ext x - exact Function.funext_iff.1 h x + exact funext_iff.1 h x @[ext] theorem ext ⦃f g : M ↪[L] N⦄ (h : ∀ x, f x = g x) : f = g := @@ -537,7 +537,7 @@ instance : EquivLike (M ≃[L] N) M N where cases g simp only [mk.injEq] ext x - exact Function.funext_iff.1 h₁ x + exact funext_iff.1 h₁ x instance : StrongHomClass L (M ≃[L] N) M N where map_fun := map_fun' diff --git a/Mathlib/ModelTheory/ElementaryMaps.lean b/Mathlib/ModelTheory/ElementaryMaps.lean index 54428842a22a5..f774228f18d50 100644 --- a/Mathlib/ModelTheory/ElementaryMaps.lean +++ b/Mathlib/ModelTheory/ElementaryMaps.lean @@ -65,7 +65,7 @@ instance instFunLike : FunLike (M ↪ₑ[L] N) M N where cases g simp only [ElementaryEmbedding.mk.injEq] ext x - exact Function.funext_iff.1 h x + exact funext_iff.1 h x instance : CoeFun (M ↪ₑ[L] N) fun _ => M → N := DFunLike.hasCoeToFun diff --git a/Mathlib/ModelTheory/Syntax.lean b/Mathlib/ModelTheory/Syntax.lean index 24c9794e5cfbd..a70ace4f16815 100644 --- a/Mathlib/ModelTheory/Syntax.lean +++ b/Mathlib/ModelTheory/Syntax.lean @@ -85,7 +85,7 @@ instance instDecidableEq [DecidableEq α] [∀ n, DecidableEq (L.Functions n)] : letI : DecidableEq (L.Term α) := instDecidableEq decidable_of_iff (f = h ▸ g ∧ ∀ i : Fin m, xs i = ys (Fin.cast h i)) <| by subst h - simp [Function.funext_iff] + simp [funext_iff] else .isFalse <| by simp [h] | .var _, .func _ _ | .func _ _, .var _ => .isFalse <| by simp diff --git a/Mathlib/NumberTheory/NumberField/House.lean b/Mathlib/NumberTheory/NumberField/House.lean index d56eb4d18c668..18b9930964456 100644 --- a/Mathlib/NumberTheory/NumberField/House.lean +++ b/Mathlib/NumberTheory/NumberField/House.lean @@ -132,7 +132,7 @@ include ha in private theorem asiegel_ne_0 : asiegel K a ≠ 0 := by simp (config := { unfoldPartialApp := true }) only [asiegel, a'] simp only [ne_eq] - rw [Function.funext_iff]; intros hs + rw [funext_iff]; intros hs simp only [Prod.forall] at hs; apply ha rw [← Matrix.ext_iff]; intros k' l @@ -141,7 +141,7 @@ private theorem asiegel_ne_0 : asiegel K a ≠ 0 := by have := ((newBasis K).repr.map_eq_zero_iff (x := (a k' l * (newBasis K) b))).1 <| by ext b' specialize hs b' - rw [Function.funext_iff] at hs + rw [funext_iff] at hs simp only [Prod.forall] at hs apply hs simp only [mul_eq_zero] at this @@ -157,7 +157,7 @@ private theorem ξ_ne_0 : ξ K x ≠ 0 := by intro H apply hxl ext ⟨l, r⟩ - rw [Function.funext_iff] at H + rw [funext_iff] at H have hblin := Basis.linearIndependent (newBasis K) simp only [zsmul_eq_mul, Fintype.linearIndependent_iff] at hblin exact hblin (fun r ↦ x (l,r)) (H _) r @@ -175,9 +175,9 @@ private theorem ξ_mulVec_eq_0 : a *ᵥ ξ K x = 0 := by have lin_0 : ∀ u, ∑ r, ∑ l, (a' K a k l r u * x (l, r) : 𝓞 K) = 0 := by intros u have hξ := ξ_ne_0 K x hxl - rw [Ne, Function.funext_iff, not_forall] at hξ + rw [Ne, funext_iff, not_forall] at hξ rcases hξ with ⟨l, hξ⟩ - rw [Function.funext_iff] at hmulvec0 + rw [funext_iff] at hmulvec0 specialize hmulvec0 ⟨k, u⟩ simp only [Fintype.sum_prod_type, mulVec, dotProduct, asiegel] at hmulvec0 rw [sum_comm] at hmulvec0 diff --git a/Mathlib/Order/RelSeries.lean b/Mathlib/Order/RelSeries.lean index d01a0e54fc8ee..b2c41395b294b 100644 --- a/Mathlib/Order/RelSeries.lean +++ b/Mathlib/Order/RelSeries.lean @@ -704,7 +704,7 @@ def injStrictMono (n : ℕ) : simp_rw [Subtype.mk_eq_mk, Sigma.mk.inj_iff, heq_eq_eq, true_and] have feq := fun i ↦ congr($(e).toFun i) simp_rw [mk_toFun lf f mf, mk_toFun lf g mg, mk_length lf f mf] at feq - rwa [Function.funext_iff] + rwa [funext_iff] /-- For two preorders `α, β`, if `f : α → β` is strictly monotonic, then a strict chain of `α` diff --git a/Mathlib/Probability/ProbabilityMassFunction/Basic.lean b/Mathlib/Probability/ProbabilityMassFunction/Basic.lean index a28a889e12f1a..14593d136558b 100644 --- a/Mathlib/Probability/ProbabilityMassFunction/Basic.lean +++ b/Mathlib/Probability/ProbabilityMassFunction/Basic.lean @@ -165,7 +165,7 @@ theorem toOuterMeasure_inj {p q : PMF α} : p.toOuterMeasure = q.toOuterMeasure theorem toOuterMeasure_apply_eq_zero_iff : p.toOuterMeasure s = 0 ↔ Disjoint p.support s := by rw [toOuterMeasure_apply, ENNReal.tsum_eq_zero] - exact Function.funext_iff.symm.trans Set.indicator_eq_zero' + exact funext_iff.symm.trans Set.indicator_eq_zero' theorem toOuterMeasure_apply_eq_one_iff : p.toOuterMeasure s = 1 ↔ p.support ⊆ s := by refine (p.toOuterMeasure_apply s).symm ▸ ⟨fun h a hap => ?_, fun h => ?_⟩ diff --git a/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean b/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean index c2cc073aa0dca..e5fead4c21a6d 100644 --- a/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean +++ b/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean @@ -98,7 +98,7 @@ def dZero : A →ₗ[k] G → A where theorem dZero_ker_eq_invariants : LinearMap.ker (dZero A) = invariants A.ρ := by ext x - simp only [LinearMap.mem_ker, mem_invariants, ← @sub_eq_zero _ _ _ x, Function.funext_iff] + simp only [LinearMap.mem_ker, mem_invariants, ← @sub_eq_zero _ _ _ x, funext_iff] rfl @[simp] theorem dZero_eq_zero [A.IsTrivial] : dZero A = 0 := by @@ -226,7 +226,7 @@ variable {A} theorem mem_oneCocycles_def (f : G → A) : f ∈ oneCocycles A ↔ ∀ g h : G, A.ρ g (f h) - f (g * h) + f g = 0 := LinearMap.mem_ker.trans <| by - rw [Function.funext_iff] + rw [funext_iff] simp only [dOne_apply, Pi.zero_apply, Prod.forall] theorem mem_oneCocycles_iff (f : G → A) : @@ -277,7 +277,7 @@ theorem mem_twoCocycles_def (f : G × G → A) : f ∈ twoCocycles A ↔ ∀ g h j : G, A.ρ g (f (h, j)) - f (g * h, j) + f (g, h * j) - f (g, h) = 0 := LinearMap.mem_ker.trans <| by - rw [Function.funext_iff] + rw [funext_iff] simp only [dTwo_apply, Prod.mk.eta, Pi.zero_apply, Prod.forall] theorem mem_twoCocycles_iff (f : G × G → A) : @@ -491,7 +491,7 @@ def twoCoboundariesOfIsTwoCoboundary {f : G × G → A} (hf : IsTwoCoboundary f) theorem isTwoCoboundary_of_twoCoboundaries (f : twoCoboundaries (Rep.ofDistribMulAction k G A)) : IsTwoCoboundary (A := A) f.1.1 := by rcases mem_range_of_mem_twoCoboundaries f.2 with ⟨x, hx⟩ - exact ⟨x, fun g h => Function.funext_iff.1 hx (g, h)⟩ + exact ⟨x, fun g h => funext_iff.1 hx (g, h)⟩ end ofDistribMulAction @@ -616,7 +616,7 @@ theorem isMulTwoCoboundary_of_twoCoboundaries (f : twoCoboundaries (Rep.ofMulDistribMulAction G M)) : IsMulTwoCoboundary (M := M) (Additive.toMul ∘ f.1.1) := by rcases mem_range_of_mem_twoCoboundaries f.2 with ⟨x, hx⟩ - exact ⟨x, fun g h => Function.funext_iff.1 hx (g, h)⟩ + exact ⟨x, fun g h => funext_iff.1 hx (g, h)⟩ end ofMulDistribMulAction diff --git a/Mathlib/RingTheory/HahnSeries/Addition.lean b/Mathlib/RingTheory/HahnSeries/Addition.lean index 61a63a5fceb4a..3daeb21892dce 100644 --- a/Mathlib/RingTheory/HahnSeries/Addition.lean +++ b/Mathlib/RingTheory/HahnSeries/Addition.lean @@ -102,7 +102,7 @@ lemma addOppositeEquiv_orderTop (x : HahnSeries Γ (Rᵃᵒᵖ)) : simp only [orderTop, AddOpposite.unop_op, mk_eq_zero, AddEquivClass.map_eq_zero_iff, addOppositeEquiv_support, ne_eq] simp only [addOppositeEquiv_apply, AddOpposite.unop_op, mk_eq_zero, zero_coeff] - simp_rw [HahnSeries.ext_iff, Function.funext_iff] + simp_rw [HahnSeries.ext_iff, funext_iff] simp only [Pi.zero_apply, AddOpposite.unop_eq_zero_iff, zero_coeff] @[simp] @@ -116,7 +116,7 @@ lemma addOppositeEquiv_leadingCoeff (x : HahnSeries Γ (Rᵃᵒᵖ)) : simp only [leadingCoeff, AddOpposite.unop_op, mk_eq_zero, AddEquivClass.map_eq_zero_iff, addOppositeEquiv_support, ne_eq] simp only [addOppositeEquiv_apply, AddOpposite.unop_op, mk_eq_zero, zero_coeff] - simp_rw [HahnSeries.ext_iff, Function.funext_iff] + simp_rw [HahnSeries.ext_iff, funext_iff] simp only [Pi.zero_apply, AddOpposite.unop_eq_zero_iff, zero_coeff] split <;> rfl diff --git a/Mathlib/RingTheory/HahnSeries/Basic.lean b/Mathlib/RingTheory/HahnSeries/Basic.lean index b063204680d7f..29586b2758fad 100644 --- a/Mathlib/RingTheory/HahnSeries/Basic.lean +++ b/Mathlib/RingTheory/HahnSeries/Basic.lean @@ -427,7 +427,7 @@ theorem embDomain_single {f : Γ ↪o Γ'} {g : Γ} {r : R} : theorem embDomain_injective {f : Γ ↪o Γ'} : Function.Injective (embDomain f : HahnSeries Γ R → HahnSeries Γ' R) := fun x y xy => by ext g - rw [HahnSeries.ext_iff, Function.funext_iff] at xy + rw [HahnSeries.ext_iff, funext_iff] at xy have xyg := xy (f g) rwa [embDomain_coeff, embDomain_coeff] at xyg diff --git a/Mathlib/RingTheory/HahnSeries/Multiplication.lean b/Mathlib/RingTheory/HahnSeries/Multiplication.lean index 153ac24993c5d..11ac4db05c9d9 100644 --- a/Mathlib/RingTheory/HahnSeries/Multiplication.lean +++ b/Mathlib/RingTheory/HahnSeries/Multiplication.lean @@ -535,7 +535,7 @@ instance instNoZeroSMulDivisors {Γ} [LinearOrderedCancelAddCommMonoid Γ] [Zero eq_zero_or_eq_zero_of_smul_eq_zero {x y} hxy := by contrapose! hxy simp only [ne_eq] - rw [HahnModule.ext_iff, Function.funext_iff, not_forall] + rw [HahnModule.ext_iff, funext_iff, not_forall] refine ⟨x.order + ((of R).symm y).order, ?_⟩ rw [smul_coeff_order_add_order x y, of_symm_zero, HahnSeries.zero_coeff, smul_eq_zero, not_or] constructor @@ -644,7 +644,7 @@ theorem C_one : C (1 : R) = (1 : HahnSeries Γ R) := theorem C_injective : Function.Injective (C : R → HahnSeries Γ R) := by intro r s rs - rw [HahnSeries.ext_iff, Function.funext_iff] at rs + rw [HahnSeries.ext_iff, funext_iff] at rs have h := rs 0 rwa [C_apply, single_coeff_same, C_apply, single_coeff_same] at h @@ -756,7 +756,7 @@ instance [Nontrivial Γ] [Nontrivial R] : Nontrivial (Subalgebra R (HahnSeries refine ⟨single a 1, ?_⟩ simp only [Algebra.mem_bot, not_exists, Set.mem_range, iff_true, Algebra.mem_top] intro x - rw [HahnSeries.ext_iff, Function.funext_iff, not_forall] + rw [HahnSeries.ext_iff, funext_iff, not_forall] refine ⟨a, ?_⟩ rw [single_coeff_same, algebraMap_apply, C_apply, single_coeff_of_ne ha] exact zero_ne_one⟩⟩ diff --git a/Mathlib/RingTheory/Ideal/Maps.lean b/Mathlib/RingTheory/Ideal/Maps.lean index 687a565ce36b9..faa5fc206c934 100644 --- a/Mathlib/RingTheory/Ideal/Maps.lean +++ b/Mathlib/RingTheory/Ideal/Maps.lean @@ -570,7 +570,7 @@ theorem ker_ne_top [Nontrivial S] (f : F) : ker f ≠ ⊤ := lemma _root_.Pi.ker_ringHom {ι : Type*} {R : ι → Type*} [∀ i, Semiring (R i)] (φ : ∀ i, S →+* R i) : ker (Pi.ringHom φ) = ⨅ i, ker (φ i) := by ext x - simp [mem_ker, Ideal.mem_iInf, Function.funext_iff] + simp [mem_ker, Ideal.mem_iInf, funext_iff] @[simp] theorem ker_rangeSRestrict (f : R →+* S) : ker f.rangeSRestrict = ker f := diff --git a/Mathlib/RingTheory/Idempotents.lean b/Mathlib/RingTheory/Idempotents.lean index c2dd1724dad6f..958157f3ad00b 100644 --- a/Mathlib/RingTheory/Idempotents.lean +++ b/Mathlib/RingTheory/Idempotents.lean @@ -402,7 +402,7 @@ lemma CompleteOrthogonalIdempotents.bijective_pi (he : CompleteOrthogonalIdempot refine ⟨?_, he.1.surjective_pi⟩ rw [injective_iff_map_eq_zero] intro x hx - simp [Function.funext_iff, Ideal.Quotient.eq_zero_iff_mem, Ideal.mem_span_singleton] at hx + simp [funext_iff, Ideal.Quotient.eq_zero_iff_mem, Ideal.mem_span_singleton] at hx suffices ∀ s : Finset I, (∏ i in s, (1 - e i)) * x = x by rw [← this Finset.univ, he.prod_one_sub, zero_mul] refine fun s ↦ Finset.induction_on s (by simp) ?_ diff --git a/Mathlib/RingTheory/WittVector/Defs.lean b/Mathlib/RingTheory/WittVector/Defs.lean index 53b4f06f4f2d8..5003404660142 100644 --- a/Mathlib/RingTheory/WittVector/Defs.lean +++ b/Mathlib/RingTheory/WittVector/Defs.lean @@ -72,7 +72,7 @@ theorem ext {x y : 𝕎 R} (h : ∀ n, x.coeff n = y.coeff n) : x = y := by cases x cases y simp only at h - simp [Function.funext_iff, h] + simp [funext_iff, h] variable (p) diff --git a/Mathlib/RingTheory/WittVector/IsPoly.lean b/Mathlib/RingTheory/WittVector/IsPoly.lean index da0cac3cda942..69617badd6694 100644 --- a/Mathlib/RingTheory/WittVector/IsPoly.lean +++ b/Mathlib/RingTheory/WittVector/IsPoly.lean @@ -113,7 +113,7 @@ theorem poly_eq_of_wittPolynomial_bind_eq' [Fact p.Prime] (f g : ℕ → MvPolyn (h : ∀ n, bind₁ f (wittPolynomial p _ n) = bind₁ g (wittPolynomial p _ n)) : f = g := by ext1 n apply MvPolynomial.map_injective (Int.castRingHom ℚ) Int.cast_injective - rw [← Function.funext_iff] at h + rw [← funext_iff] at h replace h := congr_arg (fun fam => bind₁ (MvPolynomial.map (Int.castRingHom ℚ) ∘ fam) (xInTermsOfW p ℚ n)) h simpa only [Function.comp_def, map_bind₁, map_wittPolynomial, ← bind₁_bind₁, @@ -123,7 +123,7 @@ theorem poly_eq_of_wittPolynomial_bind_eq [Fact p.Prime] (f g : ℕ → MvPolyno (h : ∀ n, bind₁ f (wittPolynomial p _ n) = bind₁ g (wittPolynomial p _ n)) : f = g := by ext1 n apply MvPolynomial.map_injective (Int.castRingHom ℚ) Int.cast_injective - rw [← Function.funext_iff] at h + rw [← funext_iff] at h replace h := congr_arg (fun fam => bind₁ (MvPolynomial.map (Int.castRingHom ℚ) ∘ fam) (xInTermsOfW p ℚ n)) h simpa only [Function.comp_def, map_bind₁, map_wittPolynomial, ← bind₁_bind₁, diff --git a/Mathlib/Topology/Algebra/Module/Basic.lean b/Mathlib/Topology/Algebra/Module/Basic.lean index 1ed9c07b215fe..3a3b76145aaeb 100644 --- a/Mathlib/Topology/Algebra/Module/Basic.lean +++ b/Mathlib/Topology/Algebra/Module/Basic.lean @@ -1134,7 +1134,7 @@ theorem pi_apply (f : ∀ i, M →L[R] φ i) (c : M) (i : ι) : pi f c i = f i c rfl theorem pi_eq_zero (f : ∀ i, M →L[R] φ i) : pi f = 0 ↔ ∀ i, f i = 0 := by - simp only [ContinuousLinearMap.ext_iff, pi_apply, Function.funext_iff] + simp only [ContinuousLinearMap.ext_iff, pi_apply, funext_iff] exact forall_swap theorem pi_zero : pi (fun _ => 0 : ∀ i, M →L[R] φ i) = 0 := diff --git a/Mathlib/Topology/Homotopy/Product.lean b/Mathlib/Topology/Homotopy/Product.lean index a898bde525861..cc848a691b504 100644 --- a/Mathlib/Topology/Homotopy/Product.lean +++ b/Mathlib/Topology/Homotopy/Product.lean @@ -72,7 +72,7 @@ def HomotopyRel.pi (homotopies : ∀ i : I, HomotopyRel (f i) (g i) S) : prop' := by intro t x hx dsimp only [coe_mk, pi_eval, toFun_eq_coe, HomotopyWith.coe_toContinuousMap] - simp only [Function.funext_iff, ← forall_and] + simp only [funext_iff, ← forall_and] intro i exact (homotopies i).prop' t x hx } diff --git a/Mathlib/Topology/LocallyConstant/Basic.lean b/Mathlib/Topology/LocallyConstant/Basic.lean index 401cf78da5d19..916e783447b5b 100644 --- a/Mathlib/Topology/LocallyConstant/Basic.lean +++ b/Mathlib/Topology/LocallyConstant/Basic.lean @@ -358,7 +358,7 @@ def unflip {X α β : Type*} [Finite α] [TopologicalSpace X] (f : α → Locall toFun x a := f a x isLocallyConstant := IsLocallyConstant.iff_isOpen_fiber.2 fun g => by have : (fun (x : X) (a : α) => f a x) ⁻¹' {g} = ⋂ a : α, f a ⁻¹' {g a} := by - ext; simp [Function.funext_iff] + ext; simp [funext_iff] rw [this] exact isOpen_iInter_of_finite fun a => (f a).isLocallyConstant _ diff --git a/Mathlib/Topology/TietzeExtension.lean b/Mathlib/Topology/TietzeExtension.lean index c707446315731..9a2271b7f1202 100644 --- a/Mathlib/Topology/TietzeExtension.lean +++ b/Mathlib/Topology/TietzeExtension.lean @@ -321,7 +321,7 @@ theorem exists_extension_forall_exists_le_ge_of_closedEmbedding [Nonempty X] (f rcases hle.eq_or_lt with (rfl | hlt) · have : ∀ x, f x = a := by simpa using hmem use const Y a - simp [this, Function.funext_iff] + simp [this, funext_iff] -- Put `c = (a + b) / 2`. Then `a < c < b` and `c - a = b - c`. set c := (a + b) / 2 have hac : a < c := left_lt_add_div_two.2 hlt diff --git a/test/instance_diamonds.lean b/test/instance_diamonds.lean index f26a6b719c7ff..8ced918dadd0d 100644 --- a/test/instance_diamonds.lean +++ b/test/instance_diamonds.lean @@ -153,7 +153,7 @@ example {k : Type _} [Semiring k] [Nontrivial k] : (Finsupp.comapSMul : SMul k (k →₀ k)) ≠ Finsupp.smulZeroClass.toSMul := by obtain ⟨u : k, hu⟩ := exists_ne (1 : k) intro h - simp only [SMul.ext_iff, @SMul.smul_eq_hSMul _ _ (_), Function.funext_iff, DFunLike.ext_iff] at h + simp only [SMul.ext_iff, @SMul.smul_eq_hSMul _ _ (_), funext_iff, DFunLike.ext_iff] at h replace h := h u (Finsupp.single 1 1) u classical rw [comapSMul_single, smul_apply, smul_eq_mul, mul_one, single_eq_same, smul_eq_mul, @@ -167,7 +167,7 @@ example {k : Type _} [Semiring k] [Nontrivial kˣ] : obtain ⟨u : kˣ, hu⟩ := exists_ne (1 : kˣ) haveI : Nontrivial k := ⟨⟨u, 1, Units.ext.ne hu⟩⟩ intro h - simp only [SMul.ext_iff, @SMul.smul_eq_hSMul _ _ (_), Function.funext_iff, DFunLike.ext_iff] at h + simp only [SMul.ext_iff, @SMul.smul_eq_hSMul _ _ (_), funext_iff, DFunLike.ext_iff] at h replace h := h u (Finsupp.single 1 1) u classical rw [comapSMul_single, smul_apply, Units.smul_def, smul_eq_mul, mul_one, single_eq_same, @@ -191,14 +191,14 @@ open Polynomial example [Semiring R] [Nontrivial R] : Polynomial.hasSMulPi _ _ ≠ (Pi.instSMul : SMul R[X] (R → R[X])) := by intro h - simp_rw [SMul.ext_iff, @SMul.smul_eq_hSMul _ _ (_), Function.funext_iff, Polynomial.ext_iff] at h + simp_rw [SMul.ext_iff, @SMul.smul_eq_hSMul _ _ (_), funext_iff, Polynomial.ext_iff] at h simpa using h X 1 1 0 /-- `Polynomial.hasSMulPi'` forms a diamond with `Pi.instSMul`. -/ example [CommSemiring R] [Nontrivial R] : Polynomial.hasSMulPi' _ _ _ ≠ (Pi.instSMul : SMul R[X] (R → R[X])) := by intro h - simp_rw [SMul.ext_iff, @SMul.smul_eq_hSMul _ _ (_), Function.funext_iff, Polynomial.ext_iff] at h + simp_rw [SMul.ext_iff, @SMul.smul_eq_hSMul _ _ (_), funext_iff, Polynomial.ext_iff] at h simpa using h X 1 1 0 -- fails `with_reducible_and_instances` #10906 From 91529560a9b0d0b1da0991450f5b1e2584ad5493 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Thu, 17 Oct 2024 05:23:32 +0000 Subject: [PATCH 047/505] chore: cleanup in Data/String/Basic (#17846) --- Mathlib/Data/String/Basic.lean | 9 ++++----- 1 file changed, 4 insertions(+), 5 deletions(-) diff --git a/Mathlib/Data/String/Basic.lean b/Mathlib/Data/String/Basic.lean index 6d86ef5ed7a82..d77a4e02a1f33 100644 --- a/Mathlib/Data/String/Basic.lean +++ b/Mathlib/Data/String/Basic.lean @@ -6,7 +6,6 @@ Authors: Mario Carneiro import Batteries.Data.String.Lemmas import Mathlib.Data.List.Lex import Mathlib.Data.Char -import Mathlib.Tactic.AdaptationNote import Mathlib.Algebra.Order.Group.Nat /-! @@ -17,6 +16,8 @@ Supplementary theorems about the `String` type. namespace String +@[simp] theorem endPos_empty : "".endPos = 0 := rfl + /-- `<` on string iterators. This coincides with `<` on strings as lists. -/ def ltb (s₁ s₂ : Iterator) : Bool := if s₂.hasNext then @@ -76,13 +77,11 @@ theorem lt_iff_toList_lt : ∀ {s₁ s₂ : String}, s₁ < s₂ ↔ s₁.toList · unfold ltb; decide · rename_i c₂ cs₂; apply iff_of_true · unfold ltb - #adaptation_note /-- v4.7.0-rc1 exclude reduceMk from simp -/ - simp [-reduceMk, Iterator.hasNext, Char.utf8Size_pos] + simp [Iterator.hasNext, Char.utf8Size_pos] · apply List.nil_lt_cons · rename_i c₁ cs₁ ih; apply iff_of_false · unfold ltb - #adaptation_note /-- v4.7.0-rc1 exclude reduceMk from simp -/ - simp [-reduceMk, Iterator.hasNext] + simp [Iterator.hasNext] · apply not_lt_of_lt; apply List.nil_lt_cons · rename_i c₁ cs₁ ih c₂ cs₂; unfold ltb simp only [Iterator.hasNext, Pos.byteIdx_zero, endPos, utf8ByteSize, utf8ByteSize.go, From 95d4d082bb8466eae74c9aacbc8f4af7f45fc320 Mon Sep 17 00:00:00 2001 From: mathlib4-update-dependencies-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Thu, 17 Oct 2024 06:03:15 +0000 Subject: [PATCH 048/505] chore: update Mathlib dependencies 2024-10-17 (#17852) This PR updates the Mathlib dependencies. --- lake-manifest.json | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lake-manifest.json b/lake-manifest.json index d4790a02b8a62..968440a4a1bb7 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,7 +5,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "422d1a5f608fccafeddab9748e8038ef346b59bf", + "rev": "d011c00010429728fab0bb3675b13f4bce63539e", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", From fff90ff17cddd549055f6f16b81b5009ef25099d Mon Sep 17 00:00:00 2001 From: Dagur Asgeirsson Date: Thu, 17 Oct 2024 06:44:29 +0000 Subject: [PATCH 049/505] feat(CategoryTheory/Limits): generalize universes in `Preserves.Filtered` and add `Reflects` classes (#17537) --- .../Limits/Preserves/Filtered.lean | 233 ++++++++++++++++-- 1 file changed, 214 insertions(+), 19 deletions(-) diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Filtered.lean b/Mathlib/CategoryTheory/Limits/Preserves/Filtered.lean index 932826e7f60e0..d1ad0cedcd3f6 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Filtered.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Filtered.lean @@ -11,8 +11,6 @@ import Mathlib.CategoryTheory.Filtered.Basic Typically forgetful functors from algebraic categories preserve filtered colimits (although not general colimits). See e.g. `Algebra/Category/MonCat/FilteredColimits`. -## Future work -This could be generalised to allow diagrams in lower universes. -/ @@ -22,7 +20,7 @@ open CategoryTheory.Functor namespace CategoryTheory.Limits -universe v u₁ u₂ u₃ +universe w' w₂' w w₂ v u₁ u₂ u₃ -- declare the `v`'s first; see `CategoryTheory.Category` for an explanation variable {C : Type u₁} [Category.{v} C] @@ -30,39 +28,236 @@ variable {D : Type u₂} [Category.{v} D] variable {E : Type u₃} [Category.{v} E] variable {J : Type v} [SmallCategory J] {K : J ⥤ C} +section FilteredColimits + +section Preserves + +-- This should be used with explicit universe variables. +/-- `PreservesFilteredColimitsOfSize.{w', w} F` means that `F` sends all colimit cocones over any +filtered diagram `J ⥤ C` to colimit cocones, where `J : Type w` with `[Category.{w'} J]`. -/ +@[nolint checkUnivs, pp_with_univ] +class PreservesFilteredColimitsOfSize (F : C ⥤ D) where + preserves_filtered_colimits : + ∀ (J : Type w) [Category.{w'} J] [IsFiltered J], PreservesColimitsOfShape J F + /-- A functor is said to preserve filtered colimits, if it preserves all colimits of shape `J`, where -`J` is a filtered category. +`J` is a filtered category which is small relative to the universe in which morphisms of the source +live. -/ -class PreservesFilteredColimits (F : C ⥤ D) : Type max u₁ u₂ (v + 1) where - preserves_filtered_colimits : - ∀ (J : Type v) [SmallCategory J] [IsFiltered J], PreservesColimitsOfShape J F +abbrev PreservesFilteredColimits (F : C ⥤ D) : Type max u₁ u₂ (v + 1) := + PreservesFilteredColimitsOfSize.{v, v} F -attribute [instance 100] PreservesFilteredColimits.preserves_filtered_colimits +attribute [instance 100] PreservesFilteredColimitsOfSize.preserves_filtered_colimits instance (priority := 100) PreservesColimits.preservesFilteredColimits (F : C ⥤ D) - [PreservesColimits F] : PreservesFilteredColimits F where + [PreservesColimitsOfSize.{w, w'} F] : PreservesFilteredColimitsOfSize.{w, w'} F where preserves_filtered_colimits _ := inferInstance -instance compPreservesFilteredColimits (F : C ⥤ D) (G : D ⥤ E) [PreservesFilteredColimits F] - [PreservesFilteredColimits G] : PreservesFilteredColimits (F ⋙ G) where +instance compPreservesFilteredColimits (F : C ⥤ D) (G : D ⥤ E) + [PreservesFilteredColimitsOfSize.{w, w'} F] [PreservesFilteredColimitsOfSize.{w, w'} G] : + PreservesFilteredColimitsOfSize.{w, w'} (F ⋙ G) where preserves_filtered_colimits _ := inferInstance -/-- A functor is said to preserve cofiltered limits, if it preserves all limits of shape `J`, where -`J` is a cofiltered category. +/-- A functor preserving larger filtered colimits also preserves smaller filtered colimits. -/ +noncomputable def preservesFilteredColimitsOfSizeOfUnivLE (F : C ⥤ D) [UnivLE.{w, w'}] + [UnivLE.{w₂, w₂'}] [PreservesFilteredColimitsOfSize.{w', w₂'} F] : + PreservesFilteredColimitsOfSize.{w, w₂} F where + preserves_filtered_colimits J _ _ := by + let e := ((ShrinkHoms.equivalence J).trans <| Shrink.equivalence _).symm + haveI := IsFiltered.of_equivalence e.symm + exact preservesColimitsOfShapeOfEquiv e F + +/-- +`PreservesFilteredColimitsOfSizeShrink.{w, w'} F` tries to obtain +`PreservesFilteredColimitsOfSize.{w, w'} F` from some other `PreservesFilteredColimitsOfSize F`. +-/ +noncomputable def preservesFilteredColimitsOfSizeShrink (F : C ⥤ D) + [PreservesFilteredColimitsOfSize.{max w w₂, max w' w₂'} F] : + PreservesFilteredColimitsOfSize.{w, w'} F := + preservesFilteredColimitsOfSizeOfUnivLE.{max w w₂, max w' w₂'} F + +/-- +Preserving filtered colimits at any universe implies preserving filtered colimits at universe `0`. +-/ +noncomputable def preservesSmallestFilteredColimitsOfPreservesFilteredColimits (F : C ⥤ D) + [PreservesFilteredColimitsOfSize.{w', w} F] : PreservesFilteredColimitsOfSize.{0, 0} F := + preservesFilteredColimitsOfSizeShrink F + +end Preserves + +section Reflects + +-- This should be used with explicit universe variables. +/-- `ReflectsFilteredColimitsOfSize.{w', w} F` means that whenever the image of a filtered cocone +under `F` is a colimit cocone, the original cocone was already a colimit. -/ +@[nolint checkUnivs, pp_with_univ] +class ReflectsFilteredColimitsOfSize (F : C ⥤ D) where + reflects_filtered_colimits : + ∀ (J : Type w) [Category.{w'} J] [IsFiltered J], ReflectsColimitsOfShape J F + +/-- +A functor is said to reflect filtered colimits, if it reflects all colimits of shape `J`, where +`J` is a filtered category which is small relative to the universe in which morphisms of the source +live. +-/ +abbrev ReflectsFilteredColimits (F : C ⥤ D) : Type max u₁ u₂ (v + 1) := + ReflectsFilteredColimitsOfSize.{v, v} F + +attribute [instance 100] ReflectsFilteredColimitsOfSize.reflects_filtered_colimits + +instance (priority := 100) ReflectsColimits.reflectsFilteredColimits (F : C ⥤ D) + [ReflectsColimitsOfSize.{w, w'} F] : ReflectsFilteredColimitsOfSize.{w, w'} F where + reflects_filtered_colimits _ := inferInstance + +instance compReflectsFilteredColimits (F : C ⥤ D) (G : D ⥤ E) + [ReflectsFilteredColimitsOfSize.{w, w'} F] [ReflectsFilteredColimitsOfSize.{w, w'} G] : + ReflectsFilteredColimitsOfSize.{w, w'} (F ⋙ G) where + reflects_filtered_colimits _ := inferInstance + +/-- A functor reflecting larger filtered colimits also reflects smaller filtered colimits. -/ +noncomputable def reflectsFilteredColimitsOfSizeOfUnivLE (F : C ⥤ D) [UnivLE.{w, w'}] + [UnivLE.{w₂, w₂'}] [ReflectsFilteredColimitsOfSize.{w', w₂'} F] : + ReflectsFilteredColimitsOfSize.{w, w₂} F where + reflects_filtered_colimits J _ _ := by + let e := ((ShrinkHoms.equivalence J).trans <| Shrink.equivalence _).symm + haveI := IsFiltered.of_equivalence e.symm + exact reflectsColimitsOfShapeOfEquiv e F + +/-- +`ReflectsFilteredColimitsOfSizeShrink.{w, w'} F` tries to obtain +`ReflectsFilteredColimitsOfSize.{w, w'} F` from some other `ReflectsFilteredColimitsOfSize F`. -/ -class PreservesCofilteredLimits (F : C ⥤ D) : Type max u₁ u₂ (v + 1) where +noncomputable def reflectsFilteredColimitsOfSizeShrink (F : C ⥤ D) + [ReflectsFilteredColimitsOfSize.{max w w₂, max w' w₂'} F] : + ReflectsFilteredColimitsOfSize.{w, w'} F := + reflectsFilteredColimitsOfSizeOfUnivLE.{max w w₂, max w' w₂'} F + +/-- +Reflecting filtered colimits at any universe implies reflecting filtered colimits at universe `0`. +-/ +noncomputable def reflectsSmallestFilteredColimitsOfReflectsFilteredColimits (F : C ⥤ D) + [ReflectsFilteredColimitsOfSize.{w', w} F] : ReflectsFilteredColimitsOfSize.{0, 0} F := + reflectsFilteredColimitsOfSizeShrink F + +end Reflects + +end FilteredColimits + +section CofilteredLimits + +section Preserves + +-- This should be used with explicit universe variables. +/-- `PreservesCofilteredLimitsOfSize.{w', w} F` means that `F` sends all limit cones over any +cofiltered diagram `J ⥤ C` to limit cones, where `J : Type w` with `[Category.{w'} J]`. -/ +@[nolint checkUnivs, pp_with_univ] +class PreservesCofilteredLimitsOfSize (F : C ⥤ D) where preserves_cofiltered_limits : - ∀ (J : Type v) [SmallCategory J] [IsCofiltered J], PreservesLimitsOfShape J F + ∀ (J : Type w) [Category.{w'} J] [IsCofiltered J], PreservesLimitsOfShape J F + +/-- +A functor is said to preserve cofiltered limits, if it preserves all limits of shape `J`, where +`J` is a cofiltered category which is small relative to the universe in which morphisms of the +source live. +-/ +abbrev PreservesCofilteredLimits (F : C ⥤ D) : Type max u₁ u₂ (v + 1) := + PreservesCofilteredLimitsOfSize.{v, v} F -attribute [instance 100] PreservesCofilteredLimits.preserves_cofiltered_limits +attribute [instance 100] PreservesCofilteredLimitsOfSize.preserves_cofiltered_limits instance (priority := 100) PreservesLimits.preservesCofilteredLimits (F : C ⥤ D) - [PreservesLimits F] : PreservesCofilteredLimits F where + [PreservesLimitsOfSize.{w, w'} F] : PreservesCofilteredLimitsOfSize.{w, w'} F where preserves_cofiltered_limits _ := inferInstance -instance compPreservesCofilteredLimits (F : C ⥤ D) (G : D ⥤ E) [PreservesCofilteredLimits F] - [PreservesCofilteredLimits G] : PreservesCofilteredLimits (F ⋙ G) where +instance compPreservesCofilteredLimits (F : C ⥤ D) (G : D ⥤ E) + [PreservesCofilteredLimitsOfSize.{w, w'} F] [PreservesCofilteredLimitsOfSize.{w, w'} G] : + PreservesCofilteredLimitsOfSize.{w, w'} (F ⋙ G) where preserves_cofiltered_limits _ := inferInstance +/-- A functor preserving larger cofiltered limits also preserves smaller cofiltered limits. -/ +noncomputable def preservesCofilteredLimitsOfSizeOfUnivLE (F : C ⥤ D) [UnivLE.{w, w'}] + [UnivLE.{w₂, w₂'}] [PreservesCofilteredLimitsOfSize.{w', w₂'} F] : + PreservesCofilteredLimitsOfSize.{w, w₂} F where + preserves_cofiltered_limits J _ _ := by + let e := ((ShrinkHoms.equivalence J).trans <| Shrink.equivalence _).symm + haveI := IsCofiltered.of_equivalence e.symm + exact preservesLimitsOfShapeOfEquiv e F + +/-- +`PreservesCofilteredLimitsOfSizeShrink.{w, w'} F` tries to obtain +`PreservesCofilteredLimitsOfSize.{w, w'} F` from some other `PreservesCofilteredLimitsOfSize F`. +-/ +noncomputable def preservesCofilteredLimitsOfSizeShrink (F : C ⥤ D) + [PreservesCofilteredLimitsOfSize.{max w w₂, max w' w₂'} F] : + PreservesCofilteredLimitsOfSize.{w, w'} F := + preservesCofilteredLimitsOfSizeOfUnivLE.{max w w₂, max w' w₂'} F + +/-- +Preserving cofiltered limits at any universe implies preserving cofiltered limits at universe `0`. +-/ +noncomputable def preservesSmallestCofilteredLimitsOfPreservesCofilteredLimits (F : C ⥤ D) + [PreservesCofilteredLimitsOfSize.{w', w} F] : PreservesCofilteredLimitsOfSize.{0, 0} F := + preservesCofilteredLimitsOfSizeShrink F + +end Preserves + +section Reflects + +-- This should be used with explicit universe variables. +/-- `ReflectsCofilteredLimitsOfSize.{w', w} F` means that whenever the image of a cofiltered cone +under `F` is a limit cone, the original cone was already a limit. -/ +@[nolint checkUnivs, pp_with_univ] +class ReflectsCofilteredLimitsOfSize (F : C ⥤ D) where + reflects_cofiltered_limits : + ∀ (J : Type w) [Category.{w'} J] [IsCofiltered J], ReflectsLimitsOfShape J F + +/-- +A functor is said to reflect cofiltered limits, if it reflects all limits of shape `J`, where +`J` is a cofiltered category which is small relative to the universe in which morphisms of the +source live. +-/ +abbrev ReflectsCofilteredLimits (F : C ⥤ D) : Type max u₁ u₂ (v + 1) := + ReflectsCofilteredLimitsOfSize.{v, v} F + +attribute [instance 100] ReflectsCofilteredLimitsOfSize.reflects_cofiltered_limits + +instance (priority := 100) ReflectsLimits.reflectsCofilteredLimits (F : C ⥤ D) + [ReflectsLimitsOfSize.{w, w'} F] : ReflectsCofilteredLimitsOfSize.{w, w'} F where + reflects_cofiltered_limits _ := inferInstance + +instance compReflectsCofilteredLimits (F : C ⥤ D) (G : D ⥤ E) + [ReflectsCofilteredLimitsOfSize.{w, w'} F] [ReflectsCofilteredLimitsOfSize.{w, w'} G] : + ReflectsCofilteredLimitsOfSize.{w, w'} (F ⋙ G) where + reflects_cofiltered_limits _ := inferInstance + +/-- A functor reflecting larger cofiltered limits also reflects smaller cofiltered limits. -/ +noncomputable def reflectsCofilteredLimitsOfSizeOfUnivLE (F : C ⥤ D) [UnivLE.{w, w'}] + [UnivLE.{w₂, w₂'}] [ReflectsCofilteredLimitsOfSize.{w', w₂'} F] : + ReflectsCofilteredLimitsOfSize.{w, w₂} F where + reflects_cofiltered_limits J _ _ := by + let e := ((ShrinkHoms.equivalence J).trans <| Shrink.equivalence _).symm + haveI := IsCofiltered.of_equivalence e.symm + exact reflectsLimitsOfShapeOfEquiv e F + +/-- +`ReflectsCofilteredLimitsOfSizeShrink.{w, w'} F` tries to obtain +`ReflectsCofilteredLimitsOfSize.{w, w'} F` from some other `ReflectsCofilteredLimitsOfSize F`. +-/ +noncomputable def reflectsCofilteredLimitsOfSizeShrink (F : C ⥤ D) + [ReflectsCofilteredLimitsOfSize.{max w w₂, max w' w₂'} F] : + ReflectsCofilteredLimitsOfSize.{w, w'} F := + reflectsCofilteredLimitsOfSizeOfUnivLE.{max w w₂, max w' w₂'} F + +/-- +Reflecting cofiltered limits at any universe implies reflecting cofiltered limits at universe `0`. +-/ +noncomputable def reflectsSmallestCofilteredLimitsOfReflectsCofilteredLimits (F : C ⥤ D) + [ReflectsCofilteredLimitsOfSize.{w', w} F] : ReflectsCofilteredLimitsOfSize.{0, 0} F := + reflectsCofilteredLimitsOfSizeShrink F + +end Reflects + +end CofilteredLimits + end CategoryTheory.Limits From b6fd8b449c9037527638f3a6a4e32943e29d27a3 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Thu, 17 Oct 2024 07:04:05 +0000 Subject: [PATCH 050/505] doc: fix some issues under RepresentationTheory (#16409) --- Mathlib/RepresentationTheory/Action/Basic.lean | 2 +- Mathlib/RepresentationTheory/Basic.lean | 13 ++++++------- Mathlib/RepresentationTheory/Character.lean | 4 ++-- .../RepresentationTheory/GroupCohomology/Basic.lean | 4 ++-- .../GroupCohomology/Resolution.lean | 6 +++--- 5 files changed, 14 insertions(+), 15 deletions(-) diff --git a/Mathlib/RepresentationTheory/Action/Basic.lean b/Mathlib/RepresentationTheory/Action/Basic.lean index 4c1ae310a5812..f70ce1fb09e29 100644 --- a/Mathlib/RepresentationTheory/Action/Basic.lean +++ b/Mathlib/RepresentationTheory/Action/Basic.lean @@ -16,7 +16,7 @@ import Mathlib.CategoryTheory.Conj The prototypical example is `V = ModuleCat R`, where `Action (ModuleCat R) G` is the category of `R`-linear representations of `G`. -We check `Action V G ≌ (singleObj G ⥤ V)`, +We check `Action V G ≌ (CategoryTheory.singleObj G ⥤ V)`, and construct the restriction functors `res {G H : MonCat} (f : G ⟶ H) : Action V H ⥤ Action V G`. -/ diff --git a/Mathlib/RepresentationTheory/Basic.lean b/Mathlib/RepresentationTheory/Basic.lean index 979151b7da507..ec67d8f8cb669 100644 --- a/Mathlib/RepresentationTheory/Basic.lean +++ b/Mathlib/RepresentationTheory/Basic.lean @@ -13,11 +13,10 @@ representations. ## Main definitions - * Representation.Representation - * Representation.character - * Representation.tprod - * Representation.linHom - * Representation.dual + * `Representation` + * `Representation.tprod` + * `Representation.linHom` + * `Representation.dual` ## Implementation notes @@ -26,7 +25,7 @@ homomorphisms `G →* (V →ₗ[k] V)`. We use the abbreviation `Representation` The theorem `asAlgebraHom_def` constructs a module over the group `k`-algebra of `G` (implemented as `MonoidAlgebra k G`) corresponding to a representation. If `ρ : Representation k G V`, this -module can be accessed via `ρ.asModule`. Conversely, given a `MonoidAlgebra k G-module `M` +module can be accessed via `ρ.asModule`. Conversely, given a `MonoidAlgebra k G`-module `M`, `M.ofModule` is the associociated representation seen as a homomorphism. -/ @@ -449,7 +448,7 @@ theorem dual_apply (g : G) : (dual ρV) g = Module.Dual.transpose (R := k) (ρV rfl /-- Given $k$-modules $V, W$, there is a homomorphism $φ : V^* ⊗ W → Hom_k(V, W)$ -(implemented by `LinearAlgebra.Contraction.dualTensorHom`). +(implemented by `dualTensorHom` in `Mathlib.LinearAlgebra.Contraction`). Given representations of $G$ on $V$ and $W$,there are representations of $G$ on $V^* ⊗ W$ and on $Hom_k(V, W)$. This lemma says that $φ$ is $G$-linear. diff --git a/Mathlib/RepresentationTheory/Character.lean b/Mathlib/RepresentationTheory/Character.lean index b68c19284a62c..5479f9417977f 100644 --- a/Mathlib/RepresentationTheory/Character.lean +++ b/Mathlib/RepresentationTheory/Character.lean @@ -19,8 +19,8 @@ is the theorem `char_orthonormal` ## Implementation notes -Irreducible representations are implemented categorically, using the `Simple` class defined in -`Mathlib.CategoryTheory.Simple` +Irreducible representations are implemented categorically, using the `CategoryTheory.Simple` class +defined in `Mathlib.CategoryTheory.Simple` ## TODO * Once we have the monoidal closed structure on `FdRep k G` and a better API for the rigid diff --git a/Mathlib/RepresentationTheory/GroupCohomology/Basic.lean b/Mathlib/RepresentationTheory/GroupCohomology/Basic.lean index 3d60e4b43bfb4..4e6c54c0981d0 100644 --- a/Mathlib/RepresentationTheory/GroupCohomology/Basic.lean +++ b/Mathlib/RepresentationTheory/GroupCohomology/Basic.lean @@ -30,8 +30,8 @@ $\mathrm{H}^n(G, A) \cong \mathrm{Ext}^n(k, A),$ where $\mathrm{Ext}$ is taken i `Rep k G`. To talk about cohomology in low degree, please see the file -`RepresentationTheory.GroupCohomology.LowDegree`, which gives simpler expressions for `H⁰, H¹, H²` -than the definition `groupCohomology` in this file. +`Mathlib.RepresentationTheory.GroupCohomology.LowDegree`, which gives simpler expressions for +`H⁰`, `H¹`, `H²` than the definition `groupCohomology` in this file. ## Main definitions diff --git a/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean b/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean index c99ba8f5afb4e..b43c932f7d3bc 100644 --- a/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean +++ b/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean @@ -23,9 +23,9 @@ This allows us to define a `k[G]`-basis on `k[Gⁿ⁺¹]`, by mapping the natura We then define the standard resolution of `k` as a trivial representation, by taking the alternating face map complex associated to an appropriate simplicial `k`-linear -`G`-representation. This simplicial object is the `linearization` of the simplicial `G`-set given -by the universal cover of the classifying space of `G`, `EG`. We prove this simplicial `G`-set `EG` -is isomorphic to the Čech nerve of the natural arrow of `G`-sets `G ⟶ {pt}`. +`G`-representation. This simplicial object is the `Rep.linearization` of the simplicial `G`-set +given by the universal cover of the classifying space of `G`, `EG`. We prove this simplicial +`G`-set `EG` is isomorphic to the Čech nerve of the natural arrow of `G`-sets `G ⟶ {pt}`. We then use this isomorphism to deduce that as a complex of `k`-modules, the standard resolution of `k` as a trivial `G`-representation is homotopy equivalent to the complex with `k` at 0 and 0 From efad68b785a481b7e8631e18eea0bc305c0414a9 Mon Sep 17 00:00:00 2001 From: Christopher Hoskin Date: Thu, 17 Oct 2024 07:12:41 +0000 Subject: [PATCH 051/505] feat(Analysis/LocallyConvex/AbsConvex): Define the Absolutely Convex Hull (#17029) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Defines the absolutely convex (or disked) hull of a subset of a locally convex space and proves a number of standard properties, including: - (subject to suitable conditions) the absolutely convex hull equals the convex hull of the balanced hull. - In a real vector space the absolutely convex hull of `s` equals the convex hull of `s ∪ -s` Co-authored-by: Christopher Hoskin --- Mathlib/Analysis/LocallyConvex/AbsConvex.lean | 192 ++++++++++++++++-- .../LocallyConvex/BalancedCoreHull.lean | 14 ++ Mathlib/Analysis/LocallyConvex/Basic.lean | 3 + 3 files changed, 197 insertions(+), 12 deletions(-) diff --git a/Mathlib/Analysis/LocallyConvex/AbsConvex.lean b/Mathlib/Analysis/LocallyConvex/AbsConvex.lean index 5cf175db643ff..f9e44391da1b4 100644 --- a/Mathlib/Analysis/LocallyConvex/AbsConvex.lean +++ b/Mathlib/Analysis/LocallyConvex/AbsConvex.lean @@ -10,23 +10,33 @@ import Mathlib.Analysis.Convex.Gauge /-! # Absolutely convex sets -A set is called absolutely convex or disked if it is convex and balanced. -The importance of absolutely convex sets comes from the fact that every locally convex +A set `s` in an commutative monoid `E` is called absolutely convex or disked if it is convex and +balanced. The importance of absolutely convex sets comes from the fact that every locally convex topological vector space has a basis consisting of absolutely convex sets. ## Main definitions +* `absConvexHull`: the absolutely convex hull of a set `s` is the smallest absolutely convex set + containing `s`. * `gaugeSeminormFamily`: the seminorm family induced by all open absolutely convex neighborhoods of zero. ## Main statements +* `absConvexHull_eq_convexHull_balancedHull`: when the locally convex space is a module, the + absolutely convex hull of a set `s` equals the convex hull of the balanced hull of `s`. +* `convexHull_union_neg_eq_absConvexHull`: the convex hull of `s ∪ -s` is the absolute convex hull + of `s`. * `with_gaugeSeminormFamily`: the topology of a locally convex space is induced by the family `gaugeSeminormFamily`. -## TODO +## Implementation notes -* Define the disked hull +Mathlib's definition of `Convex` requires the scalars to be an `OrderedSemiring` whereas the +definition of `Balanced` requires the scalars to be a `SeminormedRing`. Mathlib doesn't currently +have a concept of a semi-normed ordered ring, so we define a set as `AbsConvex` if it is balanced +over a `SeminormedRing` `𝕜` and convex over `ℝ`, assuming `IsScalarTower ℝ 𝕜 E` and +`SMulCommClass ℝ 𝕜 E` where required. ## Tags @@ -40,6 +50,104 @@ open NNReal Pointwise Topology variable {𝕜 E F G ι : Type*} +section AbsolutelyConvex + +variable (𝕜) [SeminormedRing 𝕜] [SMul 𝕜 E] [SMul ℝ E] [AddCommMonoid E] +/-- A set is absolutely convex if it is balanced and convex. Mathlib's definition of `Convex` +requires the scalars to be an `OrderedSemiring` whereas the definition of `Balanced` requires the +scalars to be a `SeminormedRing`. Mathlib doesn't currently have a concept of a semi-normed ordered +ring, so we define a set as `AbsConvex` if it is balanced over a `SeminormedRing` `𝕜` and convex +over `ℝ`. -/ +def AbsConvex (s : Set E) : Prop := Balanced 𝕜 s ∧ Convex ℝ s + +variable {𝕜} + +theorem AbsConvex.empty : AbsConvex 𝕜 (∅ : Set E) := ⟨balanced_empty, convex_empty⟩ + +theorem AbsConvex.univ : AbsConvex 𝕜 (univ : Set E) := ⟨balanced_univ, convex_univ⟩ + +theorem AbsConvex.inter {s t : Set E} (hs : AbsConvex 𝕜 s) (ht : AbsConvex 𝕜 t) : + AbsConvex 𝕜 (s ∩ t) := ⟨hs.1.inter ht.1, hs.2.inter ht.2⟩ + +theorem AbsConvex.sInter {S : Set (Set E)} (h : ∀ s ∈ S, AbsConvex 𝕜 s) : AbsConvex 𝕜 (⋂₀ S) := + ⟨.sInter fun s hs => (h s hs).1, convex_sInter fun s hs => (h s hs).2⟩ + +theorem AbsConvex.iInter {ι : Sort*} {s : ι → Set E} (h : ∀ i, AbsConvex 𝕜 (s i)) : + AbsConvex 𝕜 (⋂ i, s i) := + sInter_range s ▸ AbsConvex.sInter <| forall_mem_range.2 h + +variable (𝕜) + +/-- The absolute convex hull of a set `s` is the minimal absolute convex set that includes `s`. -/ +@[simps! isClosed] +def absConvexHull : ClosureOperator (Set E) := + .ofCompletePred (AbsConvex 𝕜) fun _ ↦ .sInter + +variable {𝕜} {s : Set E} + +theorem subset_absConvexHull : s ⊆ absConvexHull 𝕜 s := + (absConvexHull 𝕜).le_closure s + +theorem absConvex_absConvexHull : AbsConvex 𝕜 (absConvexHull 𝕜 s) := + (absConvexHull 𝕜).isClosed_closure s + +theorem balanced_absConvexHull : Balanced 𝕜 (absConvexHull 𝕜 s) := + absConvex_absConvexHull.1 + +theorem convex_absConvexHull : Convex ℝ (absConvexHull 𝕜 s) := + absConvex_absConvexHull.2 + +variable (𝕜 s) in +theorem absConvexHull_eq_iInter : + absConvexHull 𝕜 s = ⋂ (t : Set E) (_ : s ⊆ t) (_ : AbsConvex 𝕜 t), t := by + simp [absConvexHull, iInter_subtype, iInter_and] + +variable {t : Set E} {x y : E} + +theorem mem_absConvexHull_iff : x ∈ absConvexHull 𝕜 s ↔ ∀ t, s ⊆ t → AbsConvex 𝕜 t → x ∈ t := by + simp_rw [absConvexHull_eq_iInter, mem_iInter] + +theorem absConvexHull_min : s ⊆ t → AbsConvex 𝕜 t → absConvexHull 𝕜 s ⊆ t := + (absConvexHull 𝕜).closure_min + +theorem AbsConvex.absConvexHull_subset_iff (ht : AbsConvex 𝕜 t) : absConvexHull 𝕜 s ⊆ t ↔ s ⊆ t := + (show (absConvexHull 𝕜).IsClosed t from ht).closure_le_iff + +@[mono, gcongr] +theorem absConvexHull_mono (hst : s ⊆ t) : absConvexHull 𝕜 s ⊆ absConvexHull 𝕜 t := + (absConvexHull 𝕜).monotone hst + +lemma absConvexHull_eq_self : absConvexHull 𝕜 s = s ↔ AbsConvex 𝕜 s := + (absConvexHull 𝕜).isClosed_iff.symm + +alias ⟨_, AbsConvex.absConvexHull_eq⟩ := absConvexHull_eq_self + +@[simp] +theorem absConvexHull_univ : absConvexHull 𝕜 (univ : Set E) = univ := + ClosureOperator.closure_top (absConvexHull 𝕜) + +@[simp] +theorem absConvexHull_empty : absConvexHull 𝕜 (∅ : Set E) = ∅ := + AbsConvex.empty.absConvexHull_eq + +@[simp] +theorem absConvexHull_eq_empty : absConvexHull 𝕜 s = ∅ ↔ s = ∅ := by + constructor + · intro h + rw [← Set.subset_empty_iff, ← h] + exact subset_absConvexHull + · rintro rfl + exact absConvexHull_empty + +@[simp] +theorem absConvexHull_nonempty : (absConvexHull 𝕜 s).Nonempty ↔ s.Nonempty := by + rw [nonempty_iff_ne_empty, nonempty_iff_ne_empty, Ne, Ne] + exact not_congr absConvexHull_eq_empty + +protected alias ⟨_, Set.Nonempty.absConvexHull⟩ := absConvexHull_nonempty + +end AbsolutelyConvex + section NontriviallyNormedField variable (𝕜 E) {s : Set E} @@ -47,8 +155,8 @@ variable [NontriviallyNormedField 𝕜] [AddCommGroup E] [Module 𝕜 E] variable [Module ℝ E] [SMulCommClass ℝ 𝕜 E] variable [TopologicalSpace E] [LocallyConvexSpace ℝ E] [ContinuousSMul 𝕜 E] -theorem nhds_basis_abs_convex : - (𝓝 (0 : E)).HasBasis (fun s : Set E => s ∈ 𝓝 (0 : E) ∧ Balanced 𝕜 s ∧ Convex ℝ s) id := by +theorem nhds_hasBasis_absConvex : + (𝓝 (0 : E)).HasBasis (fun s : Set E => s ∈ 𝓝 (0 : E) ∧ AbsConvex 𝕜 s) id := by refine (LocallyConvexSpace.convex_basis_zero ℝ E).to_hasBasis (fun s hs => ?_) fun s hs => ⟨s, ⟨hs.1, hs.2.2⟩, rfl.subset⟩ @@ -59,9 +167,9 @@ theorem nhds_basis_abs_convex : variable [ContinuousSMul ℝ E] [TopologicalAddGroup E] -theorem nhds_basis_abs_convex_open : - (𝓝 (0 : E)).HasBasis (fun s => (0 : E) ∈ s ∧ IsOpen s ∧ Balanced 𝕜 s ∧ Convex ℝ s) id := by - refine (nhds_basis_abs_convex 𝕜 E).to_hasBasis ?_ ?_ +theorem nhds_hasBasis_absConvex_open : + (𝓝 (0 : E)).HasBasis (fun s => (0 : E) ∈ s ∧ IsOpen s ∧ AbsConvex 𝕜 s) id := by + refine (nhds_hasBasis_absConvex 𝕜 E).to_hasBasis ?_ ?_ · rintro s ⟨hs_nhds, hs_balanced, hs_convex⟩ refine ⟨interior s, ?_, interior_subset⟩ exact @@ -72,6 +180,66 @@ theorem nhds_basis_abs_convex_open : end NontriviallyNormedField +section + +variable (𝕜) [NontriviallyNormedField 𝕜] +variable [AddCommGroup E] [Module ℝ E] [Module 𝕜 E] + +theorem absConvexHull_add_subset {s t : Set E} : + absConvexHull 𝕜 (s + t) ⊆ absConvexHull 𝕜 s + absConvexHull 𝕜 t := + absConvexHull_min (add_subset_add subset_absConvexHull subset_absConvexHull) + ⟨Balanced.add balanced_absConvexHull balanced_absConvexHull, + Convex.add convex_absConvexHull convex_absConvexHull⟩ + +theorem absConvexHull_eq_convexHull_balancedHull [SMulCommClass ℝ 𝕜 E] {s : Set E} : + absConvexHull 𝕜 s = convexHull ℝ (balancedHull 𝕜 s) := le_antisymm + (absConvexHull_min + ((subset_convexHull ℝ s).trans (convexHull_mono (subset_balancedHull 𝕜))) + ⟨Balanced.convexHull (balancedHull.balanced s), convex_convexHull ..⟩) + (convexHull_min (balanced_absConvexHull.balancedHull_subset_of_subset subset_absConvexHull) + convex_absConvexHull) + +/-- In general, equality doesn't hold here - e.g. consider `s := {(-1, 1), (1, 1)}` in `ℝ²`. -/ +theorem balancedHull_convexHull_subseteq_absConvexHull {s : Set E} : + balancedHull 𝕜 (convexHull ℝ s) ⊆ absConvexHull 𝕜 s := + balanced_absConvexHull.balancedHull_subset_of_subset + (convexHull_min subset_absConvexHull convex_absConvexHull) + +end + +section + +variable [AddCommGroup E] [Module ℝ E] + +lemma balancedHull_subset_convexHull_union_neg {s : Set E} : + balancedHull ℝ s ⊆ convexHull ℝ (s ∪ -s) := by + intro a ha + obtain ⟨r, hr, y, hy, rfl⟩ := mem_balancedHull_iff.1 ha + apply segment_subset_convexHull (mem_union_left (-s) hy) (mem_union_right _ (neg_mem_neg.mpr hy)) + refine ⟨(1 + r)/2, (1 - r)/2, ?_, ?_⟩ + · rw [← zero_div 2] + exact (div_le_div_right zero_lt_two).mpr (neg_le_iff_add_nonneg'.mp (neg_le_of_abs_le hr)) + · constructor + · rw [← zero_div 2] + exact (div_le_div_right zero_lt_two).mpr (sub_nonneg_of_le (le_of_max_le_left hr)) + · constructor + · ring_nf + · rw [smul_neg, ← sub_eq_add_neg, ← sub_smul] + apply congrFun (congrArg HSMul.hSMul _) y + ring_nf + +@[simp] +theorem convexHull_union_neg_eq_absConvexHull {s : Set E} : + convexHull ℝ (s ∪ -s) = absConvexHull ℝ s := by + rw [absConvexHull_eq_convexHull_balancedHull] + exact le_antisymm (convexHull_mono (union_subset (subset_balancedHull ℝ) + (fun _ _ => by rw [mem_balancedHull_iff]; use -1; aesop))) + (by + rw [← Convex.convexHull_eq (convex_convexHull ℝ (s ∪ -s))] + exact convexHull_mono balancedHull_subset_convexHull_union_neg) + +end + section AbsolutelyConvexSets variable [TopologicalSpace E] [AddCommMonoid E] [Zero E] [SeminormedRing 𝕜] @@ -80,7 +248,7 @@ variable (𝕜 E) /-- The type of absolutely convex open sets. -/ def AbsConvexOpenSets := - { s : Set E // (0 : E) ∈ s ∧ IsOpen s ∧ Balanced 𝕜 s ∧ Convex ℝ s } + { s : Set E // (0 : E) ∈ s ∧ IsOpen s ∧ AbsConvex 𝕜 s } noncomputable instance AbsConvexOpenSets.instCoeTC : CoeTC (AbsConvexOpenSets 𝕜 E) (Set E) := ⟨Subtype.val⟩ @@ -139,7 +307,7 @@ variable [SMulCommClass ℝ 𝕜 E] [LocallyConvexSpace ℝ E] /-- The topology of a locally convex space is induced by the gauge seminorm family. -/ theorem with_gaugeSeminormFamily : WithSeminorms (gaugeSeminormFamily 𝕜 E) := by refine SeminormFamily.withSeminorms_of_hasBasis _ ?_ - refine (nhds_basis_abs_convex_open 𝕜 E).to_hasBasis (fun s hs => ?_) fun s hs => ?_ + refine (nhds_hasBasis_absConvex_open 𝕜 E).to_hasBasis (fun s hs => ?_) fun s hs => ?_ · refine ⟨s, ⟨?_, rfl.subset⟩⟩ convert (gaugeSeminormFamily _ _).basisSets_singleton_mem ⟨s, hs⟩ one_pos rw [gaugeSeminormFamily_ball, Subtype.coe_mk] @@ -152,7 +320,7 @@ theorem with_gaugeSeminormFamily : WithSeminorms (gaugeSeminormFamily 𝕜 E) := ⟨mem_iInter₂.mpr fun _ _ => by simp [Seminorm.mem_ball_zero, hr], isOpen_biInter_finset fun S _ => ?_, balanced_iInter₂ fun _ _ => Seminorm.balanced_ball_zero _ _, - convex_iInter₂ fun _ _ => Seminorm.convex_ball _ _ _⟩ + convex_iInter₂ fun _ _ => Seminorm.convex_ball ..⟩ -- The only nontrivial part is to show that the ball is open have hr' : r = ‖(r : 𝕜)‖ * 1 := by simp [abs_of_pos hr] have hr'' : (r : 𝕜) ≠ 0 := by simp [hr.ne'] diff --git a/Mathlib/Analysis/LocallyConvex/BalancedCoreHull.lean b/Mathlib/Analysis/LocallyConvex/BalancedCoreHull.lean index 08b82218840f4..bf2b65c776bf0 100644 --- a/Mathlib/Analysis/LocallyConvex/BalancedCoreHull.lean +++ b/Mathlib/Analysis/LocallyConvex/BalancedCoreHull.lean @@ -104,6 +104,14 @@ theorem Balanced.balancedHull_subset_of_subset (ht : Balanced 𝕜 t) (h : s ⊆ obtain ⟨r, hr, y, hy, rfl⟩ := mem_balancedHull_iff.1 hx exact ht.smul_mem hr (h hy) +@[mono, gcongr] +theorem balancedHull_mono (hst : s ⊆ t) : balancedHull 𝕜 s ⊆ balancedHull 𝕜 t := by + intro x hx + rw [mem_balancedHull_iff] at * + obtain ⟨r, hr₁, hr₂⟩ := hx + use r + exact ⟨hr₁, smul_set_mono hst hr₂⟩ + end SMul section Module @@ -133,6 +141,12 @@ theorem balancedHull.balanced (s : Set E) : Balanced 𝕜 (balancedHull 𝕜 s) rw [← smul_assoc] at hx exact ⟨a • r, (SeminormedRing.norm_mul _ _).trans (mul_le_one₀ ha (norm_nonneg r) hr), hx⟩ +open Balanced in +theorem balancedHull_add_subset [NormOneClass 𝕜] {t : Set E} : + balancedHull 𝕜 (s + t) ⊆ balancedHull 𝕜 s + balancedHull 𝕜 t := + balancedHull_subset_of_subset (add (balancedHull.balanced _) (balancedHull.balanced _)) + (add_subset_add (subset_balancedHull _) (subset_balancedHull _)) + end Module end SeminormedRing diff --git a/Mathlib/Analysis/LocallyConvex/Basic.lean b/Mathlib/Analysis/LocallyConvex/Basic.lean index b8e661ddf003f..ffe0e4d997485 100644 --- a/Mathlib/Analysis/LocallyConvex/Basic.lean +++ b/Mathlib/Analysis/LocallyConvex/Basic.lean @@ -95,6 +95,9 @@ theorem balanced_iUnion₂ {f : ∀ i, κ i → Set E} (h : ∀ i j, Balanced Balanced 𝕜 (⋃ (i) (j), f i j) := balanced_iUnion fun _ => balanced_iUnion <| h _ +theorem Balanced.sInter {S : Set (Set E)} (h : ∀ s ∈ S, Balanced 𝕜 s) : Balanced 𝕜 (⋂₀ S) := + fun _ _ => (smul_set_sInter_subset ..).trans (fun _ _ => by aesop) + theorem balanced_iInter {f : ι → Set E} (h : ∀ i, Balanced 𝕜 (f i)) : Balanced 𝕜 (⋂ i, f i) := fun _a ha => (smul_set_iInter_subset _ _).trans <| iInter_mono fun _ => h _ _ ha From 4f4f4ec65ea3234d6eb302a3356ef2376cc5afa3 Mon Sep 17 00:00:00 2001 From: Hannah Fechtner <123216642+hannahfechtner@users.noreply.github.com> Date: Thu, 17 Oct 2024 07:12:43 +0000 Subject: [PATCH 052/505] feat(Algebra/FreeMonoid/Basic): define length, symbols, membership, reversing, induction principles (#17180) add these properties for future use in the braid monoid (I have avoided duplicating the List API when possible, ex: using FreeMonoid.lift. These are the remaining properties I need to add) Co-authored-by: Hannah Fechtner --- Mathlib/Algebra/FreeMonoid/Basic.lean | 147 ++++++++++++++++++ .../Algebra/Group/Submonoid/Membership.lean | 6 +- Mathlib/Control/Fold.lean | 7 +- Mathlib/GroupTheory/Coprod/Basic.lean | 12 +- 4 files changed, 160 insertions(+), 12 deletions(-) diff --git a/Mathlib/Algebra/FreeMonoid/Basic.lean b/Mathlib/Algebra/FreeMonoid/Basic.lean index c2a87d24b37d2..3ad679d33b048 100644 --- a/Mathlib/Algebra/FreeMonoid/Basic.lean +++ b/Mathlib/Algebra/FreeMonoid/Basic.lean @@ -106,6 +106,69 @@ theorem toList_of_mul (x : α) (xs : FreeMonoid α) : toList (of x * xs) = x :: @[to_additive] theorem of_injective : Function.Injective (@of α) := List.singleton_injective +/-! ### Length -/ + +section Length +variable {a : FreeMonoid α} + +/-- The length of a free monoid element: 1.length = 0 and (a * b).length = a.length + b.length -/ +@[to_additive "The length of an additive free monoid element: 1.length = 0 and (a + b).length = + a.length + b.length"] +def length (a : FreeMonoid α) : ℕ := List.length a + +@[to_additive (attr := simp)] +theorem length_one : length (1 : FreeMonoid α) = 0 := rfl + +@[to_additive (attr := simp)] +theorem length_eq_zero : length a = 0 ↔ a = 1 := List.length_eq_zero + +@[to_additive (attr := simp)] +theorem length_of (m : α) : length (of m) = 1 := rfl + +@[to_additive existing] +theorem length_eq_one : length a = 1 ↔ ∃ m, a = FreeMonoid.of m := + List.length_eq_one + +@[to_additive] +theorem length_eq_two {v : FreeMonoid α} : + v.length = 2 ↔ ∃ c d, v = FreeMonoid.of c * FreeMonoid.of d := List.length_eq_two + +@[to_additive (attr := simp)] +theorem length_mul (a b : FreeMonoid α) : (a * b).length = a.length + b.length := + List.length_append _ _ + +@[to_additive] +theorem of_ne_one (a : α) : of a ≠ 1 := by + intro h + have := congrArg FreeMonoid.length h + simp only [length_of, length_one, Nat.succ_ne_self] at this + +end Length + +section Mem +variable {m : α} + +/-- Membership in a free monoid element -/ +@[to_additive "Membership in a free monoid element"] +def mem (a : FreeMonoid α) (m : α) := m ∈ toList a + +@[to_additive] +instance : Membership α (FreeMonoid α) := ⟨mem⟩ + +@[to_additive] +theorem not_mem_one : ¬ m ∈ (1 : FreeMonoid α) := List.not_mem_nil _ + +@[to_additive (attr := simp)] +theorem mem_of {n : α} : m ∈ of n ↔ m = n := List.mem_singleton + +@[to_additive] +theorem mem_of_self : m ∈ of m := List.mem_singleton_self _ + +@[to_additive (attr := simp)] +theorem mem_mul {a b : FreeMonoid α} : m ∈ (a * b) ↔ m ∈ a ∨ m ∈ b := List.mem_append + +end Mem + /-- Recursor for `FreeMonoid` using `1` and `FreeMonoid.of x * xs` instead of `[]` and `x :: xs`. -/ @[to_additive (attr := elab_as_elim, induction_eliminator) "Recursor for `FreeAddMonoid` using `0` and @@ -123,6 +186,27 @@ theorem recOn_of_mul {C : FreeMonoid α → Sort*} (x : α) (xs : FreeMonoid α) (ih : ∀ x xs, C xs → C (of x * xs)) : @recOn α C (of x * xs) h0 ih = ih x xs (recOn xs h0 ih) := rfl +/-! ### Induction -/ + +section induction_principles + +/-- An induction principle on free monoids, with cases for `1`, `FreeMonoid.of` and `*`. -/ +@[to_additive (attr := elab_as_elim, induction_eliminator) +"An induction principle on free monoids, with cases for `0`, `FreeAddMonoid.of` and `+`."] +protected theorem inductionOn {C : FreeMonoid α → Prop} (z : FreeMonoid α) (one : C 1) + (of : ∀ (x : α), C (FreeMonoid.of x)) (mul : ∀ (x y : FreeMonoid α), C x → C y → C (x * y)) : + C z := List.rec one (fun _ _ ih => mul [_] _ (of _) ih) z + +/-- An induction principle for free monoids which mirrors induction on lists, with cases analogous +to the empty list and cons -/ +@[to_additive (attr := elab_as_elim) "An induction principle for free monoids which mirrors +induction on lists, with cases analogous to the empty list and cons"] +protected theorem inductionOn' {p : FreeMonoid α → Prop} (a : FreeMonoid α) + (one : p (1 : FreeMonoid α)) (mul_of : ∀ b a, p a → p (of b * a)) : p a := + List.rec one (fun _ _ tail_ih => mul_of _ _ tail_ih) a + +end induction_principles + /-- A version of `List.cases_on` for `FreeMonoid` using `1` and `FreeMonoid.of x * xs` instead of `[]` and `x :: xs`. -/ @[to_additive (attr := elab_as_elim, cases_eliminator) @@ -220,6 +304,10 @@ theorem of_smul (f : α → β → β) (x : α) (y : β) : (haveI := mkMulAction f of x • y) = f x y := rfl +/-! ### map -/ + +section Map +variable {f : α → β} {a b : FreeMonoid α} /-- The unique monoid homomorphism `FreeMonoid α →* FreeMonoid β` that sends each `of x` to `of (f x)`. -/ @[to_additive "The unique additive monoid homomorphism `FreeAddMonoid α →+ FreeAddMonoid β` @@ -232,6 +320,15 @@ def map (f : α → β) : FreeMonoid α →* FreeMonoid β where @[to_additive (attr := simp)] theorem map_of (f : α → β) (x : α) : map f (of x) = of (f x) := rfl +@[to_additive] +theorem mem_map {m : β} : m ∈ map f a ↔ ∃ n ∈ a, f n = m := List.mem_map + +@[to_additive] +theorem map_map {α₁ : Type*} {g : α₁ → α} {x : FreeMonoid α₁} : + map f (map g x) = map (f ∘ g) x := by + unfold map + simp only [MonoidHom.coe_mk, OneHom.coe_mk, toList_ofList, List.map_map] + @[to_additive] theorem toList_map (f : α → β) (xs : FreeMonoid α) : toList (map f xs) = xs.toList.map f := rfl @@ -254,4 +351,54 @@ instance uniqueUnits : Unique (FreeMonoid α)ˣ where have : toList u.val ++ toList u.inv = [] := DFunLike.congr_arg toList u.val_inv (List.append_eq_nil.mp this).1 +@[to_additive (attr := simp)] +theorem map_surjective {f : α → β} : Function.Surjective (map f) ↔ Function.Surjective f := by + constructor + · intro fs d + rcases fs (FreeMonoid.of d) with ⟨b, hb⟩ + induction' b using FreeMonoid.inductionOn' with head _ _ + · have H := congr_arg length hb + simp only [length_one, length_of, Nat.zero_ne_one, map_one] at H + simp only [map_mul, map_of] at hb + use head + have H := congr_arg length hb + simp only [length_mul, length_of, add_right_eq_self, length_eq_zero] at H + rw [H, mul_one] at hb + exact FreeMonoid.of_injective hb + intro fs d + induction' d using FreeMonoid.inductionOn' with head tail ih + · use 1 + rfl + specialize fs head + rcases fs with ⟨a, rfl⟩ + rcases ih with ⟨b, rfl⟩ + use FreeMonoid.of a * b + rfl + +end Map + +/-! ### reverse -/ + +section Reverse +/-- reverses the symbols in a free monoid element -/ +@[to_additive "reverses the symbols in an additive free monoid element"] +def reverse : FreeMonoid α → FreeMonoid α := List.reverse + +@[to_additive (attr := simp)] +theorem reverse_of (a : α) : reverse (of a) = of a := rfl + +@[to_additive] +theorem reverse_mul {a b : FreeMonoid α} : reverse (a * b) = reverse b * reverse a := + List.reverse_append _ _ + +@[to_additive (attr := simp)] +theorem reverse_reverse {a : FreeMonoid α} : reverse (reverse a) = a := by + apply List.reverse_reverse + +@[to_additive (attr := simp)] +theorem length_reverse {a : FreeMonoid α} : a.reverse.length = a.length := + List.length_reverse _ + +end Reverse + end FreeMonoid diff --git a/Mathlib/Algebra/Group/Submonoid/Membership.lean b/Mathlib/Algebra/Group/Submonoid/Membership.lean index 632e86eea4ba5..2c5e76300f647 100644 --- a/Mathlib/Algebra/Group/Submonoid/Membership.lean +++ b/Mathlib/Algebra/Group/Submonoid/Membership.lean @@ -356,9 +356,9 @@ theorem closure_induction_left {s : Set M} {p : (m : M) → m ∈ closure s → p x h := by simp_rw [closure_eq_mrange] at h obtain ⟨l, rfl⟩ := h - induction l with - | h0 => exact one - | ih x y ih => + induction l using FreeMonoid.inductionOn' with + | one => exact one + | mul_of x y ih => simp only [map_mul, FreeMonoid.lift_eval_of] refine mul_left _ x.prop (FreeMonoid.lift Subtype.val y) _ (ih ?_) simp only [closure_eq_mrange, mem_mrange, exists_apply_eq_apply] diff --git a/Mathlib/Control/Fold.lean b/Mathlib/Control/Fold.lean index 711e71a910635..62e4e80dddadb 100644 --- a/Mathlib/Control/Fold.lean +++ b/Mathlib/Control/Fold.lean @@ -301,9 +301,10 @@ theorem toList_spec (xs : t α) : toList xs = FreeMonoid.toList (foldMap FreeMon calc FreeMonoid.toList (foldMap FreeMonoid.of xs) = FreeMonoid.toList (foldMap FreeMonoid.of xs).reverse.reverse := by - simp only [List.reverse_reverse] - _ = FreeMonoid.toList (List.foldr cons [] (foldMap FreeMonoid.of xs).reverse).reverse := by - simp only [List.foldr_eta] + simp only [FreeMonoid.reverse_reverse] + _ = (List.foldr cons [] (foldMap FreeMonoid.of xs).toList.reverse).reverse := by + simp only [FreeMonoid.reverse_reverse, List.foldr_reverse, List.foldl_flip_cons_eq_append, + List.append_nil, List.reverse_reverse] _ = (unop (Foldl.ofFreeMonoid (flip cons) (foldMap FreeMonoid.of xs)) []).reverse := by #adaptation_note /-- nightly-2024-03-16: simp was simp [flip, List.foldr_reverse, Foldl.ofFreeMonoid, unop_op] -/ diff --git a/Mathlib/GroupTheory/Coprod/Basic.lean b/Mathlib/GroupTheory/Coprod/Basic.lean index 2349d4a1dad9b..169ff1d0e0bc6 100644 --- a/Mathlib/GroupTheory/Coprod/Basic.lean +++ b/Mathlib/GroupTheory/Coprod/Basic.lean @@ -191,9 +191,9 @@ theorem induction_on' {C : M ∗ N → Prop} (m : M ∗ N) (inl_mul : ∀ m x, C x → C (inl m * x)) (inr_mul : ∀ n x, C x → C (inr n * x)) : C m := by rcases mk_surjective m with ⟨x, rfl⟩ - induction x with - | h0 => exact one - | ih x xs ih => + induction x using FreeMonoid.inductionOn' with + | one => exact one + | mul_of x xs ih => cases x with | inl m => simpa using inl_mul m _ ih | inr n => simpa using inr_mul n _ ih @@ -573,9 +573,9 @@ theorem mk_of_inv_mul : ∀ x : G ⊕ H, mk (of (x.map Inv.inv Inv.inv)) * mk (o theorem con_inv_mul_cancel (x : FreeMonoid (G ⊕ H)) : coprodCon G H (ofList (x.toList.map (Sum.map Inv.inv Inv.inv)).reverse * x) 1 := by rw [← mk_eq_mk, map_mul, map_one] - induction x with - | h0 => simp [map_one mk] -- TODO: fails without `[map_one mk]` - | ih x xs ihx => + induction x using FreeMonoid.inductionOn' with + | one => simp + | mul_of x xs ihx => simp only [toList_of_mul, map_cons, reverse_cons, ofList_append, map_mul, ihx, ofList_singleton] rwa [mul_assoc, ← mul_assoc (mk (of _)), mk_of_inv_mul, one_mul] From dc48f47da5a6275c7d25a00cf82c2684af1e754d Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Thu, 17 Oct 2024 07:12:44 +0000 Subject: [PATCH 053/505] feat(CategoryTheory/Products): an ext lemma for morphisms in product categories (#17721) Until now, in order to check an identity of morphisms in a product category, we had to do `dsimp; ext`. Now, `ext` would also work, which shall ease automation in some future applications. --- Mathlib/CategoryTheory/Abelian/NonPreadditive.lean | 4 ++-- Mathlib/CategoryTheory/Closed/Ideal.lean | 2 +- Mathlib/CategoryTheory/Generator.lean | 2 +- Mathlib/CategoryTheory/Limits/Connected.lean | 4 ++-- .../CategoryTheory/Limits/Shapes/NormalMono/Equalizers.lean | 4 ++-- Mathlib/CategoryTheory/Monad/Products.lean | 2 +- Mathlib/CategoryTheory/Monoidal/OfHasFiniteProducts.lean | 2 +- Mathlib/CategoryTheory/Products/Basic.lean | 5 +++++ 8 files changed, 15 insertions(+), 10 deletions(-) diff --git a/Mathlib/CategoryTheory/Abelian/NonPreadditive.lean b/Mathlib/CategoryTheory/Abelian/NonPreadditive.lean index 680956963fe88..810dadb3c8117 100644 --- a/Mathlib/CategoryTheory/Abelian/NonPreadditive.lean +++ b/Mathlib/CategoryTheory/Abelian/NonPreadditive.lean @@ -232,12 +232,12 @@ instance epi_r {A : C} : Epi (r A) := by let hp1 : IsLimit (KernelFork.ofι (prod.lift (𝟙 A) (0 : A ⟶ A)) hlp) := by refine Fork.IsLimit.mk _ (fun s => Fork.ι s ≫ Limits.prod.fst) ?_ ?_ · intro s - apply prod.hom_ext <;> simp + apply Limits.prod.hom_ext <;> simp · intro s m h haveI : Mono (prod.lift (𝟙 A) (0 : A ⟶ A)) := mono_of_mono_fac (prod.lift_fst _ _) apply (cancel_mono (prod.lift (𝟙 A) (0 : A ⟶ A))).1 convert h - apply prod.hom_ext <;> simp + apply Limits.prod.hom_ext <;> simp let hp2 : IsColimit (CokernelCofork.ofπ (Limits.prod.snd : A ⨯ A ⟶ A) hlp) := epiIsCokernelOfKernel _ hp1 apply NormalMonoCategory.epi_of_zero_cancel diff --git a/Mathlib/CategoryTheory/Closed/Ideal.lean b/Mathlib/CategoryTheory/Closed/Ideal.lean index da5c0efda1a01..588c491499a4f 100644 --- a/Mathlib/CategoryTheory/Closed/Ideal.lean +++ b/Mathlib/CategoryTheory/Closed/Ideal.lean @@ -204,7 +204,7 @@ theorem bijection_symm_apply_id (A B : C) : rw [prod.comp_lift_assoc, prod.lift_snd, prod.lift_fst_assoc, prod.lift_fst_comp_snd_comp, ← Adjunction.eq_unit_comp_map_iff, Iso.comp_inv_eq, assoc] rw [PreservesLimitPair.iso_hom i ((reflector i).obj A) ((reflector i).obj B)] - apply prod.hom_ext + apply Limits.prod.hom_ext · rw [Limits.prod.map_fst, assoc, assoc, prodComparison_fst, ← i.map_comp, prodComparison_fst] apply (reflectorAdjunction i).unit.naturality · rw [Limits.prod.map_snd, assoc, assoc, prodComparison_snd, ← i.map_comp, prodComparison_snd] diff --git a/Mathlib/CategoryTheory/Generator.lean b/Mathlib/CategoryTheory/Generator.lean index 0424d4339d386..91d5fd9417252 100644 --- a/Mathlib/CategoryTheory/Generator.lean +++ b/Mathlib/CategoryTheory/Generator.lean @@ -509,7 +509,7 @@ theorem isCoseparator_prod (G H : C) [HasBinaryProduct G H] : refine ⟨fun h X Y u v huv => ?_, fun h => (isCoseparator_def _).2 fun X Y u v huv => h _ _ fun Z hZ g => ?_⟩ - · refine h.def _ _ fun g => prod.hom_ext ?_ ?_ + · refine h.def _ _ fun g => Limits.prod.hom_ext ?_ ?_ · simpa using huv G (by simp) (g ≫ Limits.prod.fst) · simpa using huv H (by simp) (g ≫ Limits.prod.snd) · simp only [Set.mem_insert_iff, Set.mem_singleton_iff] at hZ diff --git a/Mathlib/CategoryTheory/Limits/Connected.lean b/Mathlib/CategoryTheory/Limits/Connected.lean index 82a4900695c0d..712b12fbf1148 100644 --- a/Mathlib/CategoryTheory/Limits/Connected.lean +++ b/Mathlib/CategoryTheory/Limits/Connected.lean @@ -95,12 +95,12 @@ noncomputable def prodPreservesConnectedLimits [IsConnected J] (X : C) : { lift := fun s => prod.lift (s.π.app (Classical.arbitrary _) ≫ Limits.prod.fst) (l.lift (forgetCone s)) fac := fun s j => by - apply prod.hom_ext + apply Limits.prod.hom_ext · erw [assoc, limMap_π, comp_id, limit.lift_π] exact (nat_trans_from_is_connected (s.π ≫ γ₁ X) j (Classical.arbitrary _)).symm · simp [← l.fac (forgetCone s) j] uniq := fun s m L => by - apply prod.hom_ext + apply Limits.prod.hom_ext · erw [limit.lift_π, ← L (Classical.arbitrary J), assoc, limMap_π, comp_id] rfl · rw [limit.lift_π] diff --git a/Mathlib/CategoryTheory/Limits/Shapes/NormalMono/Equalizers.lean b/Mathlib/CategoryTheory/Limits/Shapes/NormalMono/Equalizers.lean index c71d6af56f0d8..67d189d906d5d 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/NormalMono/Equalizers.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/NormalMono/Equalizers.lean @@ -55,7 +55,7 @@ def pullback_of_mono {X Y Z : C} (a : X ⟶ Z) (b : Y ⟶ Z) [Mono a] [Mono b] : PullbackCone.IsLimit.mk _ (fun s => kernel.lift (prod.lift f g) (PullbackCone.snd s ≫ b) <| - prod.hom_ext + Limits.prod.hom_ext (calc ((PullbackCone.snd s ≫ b) ≫ prod.lift f g) ≫ Limits.prod.fst = PullbackCone.snd s ≫ b ≫ f := by simp only [prod.lift_fst, Category.assoc] @@ -125,7 +125,7 @@ def hasLimit_parallelPair {X Y : C} (f g : X ⟶ Y) : HasLimit (parallelPair f g Fork.IsLimit.mk _ (fun s => pullback.lift (Fork.ι s) (Fork.ι s) <| - prod.hom_ext (by simp only [prod.lift_fst, Category.assoc]) + Limits.prod.hom_ext (by simp only [prod.lift_fst, Category.assoc]) (by simp only [prod.comp_lift, Fork.condition s])) (fun s => by simp) fun s m h => pullback.hom_ext (by simpa only [pullback.lift_fst] using h) diff --git a/Mathlib/CategoryTheory/Monad/Products.lean b/Mathlib/CategoryTheory/Monad/Products.lean index 2a42c1d2a6ff1..16ba3d6f08034 100644 --- a/Mathlib/CategoryTheory/Monad/Products.lean +++ b/Mathlib/CategoryTheory/Monad/Products.lean @@ -73,7 +73,7 @@ def coalgebraEquivOver : Coalgebra (prodComonad X) ≌ Over X where functor := coalgebraToOver X inverse := overToCoalgebra X unitIso := NatIso.ofComponents fun A => - Coalgebra.isoMk (Iso.refl _) (prod.hom_ext (by simp) (by simpa using A.counit)) + Coalgebra.isoMk (Iso.refl _) (Limits.prod.hom_ext (by simp) (by simpa using A.counit)) counitIso := NatIso.ofComponents fun f => Over.isoMk (Iso.refl _) end diff --git a/Mathlib/CategoryTheory/Monoidal/OfHasFiniteProducts.lean b/Mathlib/CategoryTheory/Monoidal/OfHasFiniteProducts.lean index 09ea239996ebd..9e49dda0d93bf 100644 --- a/Mathlib/CategoryTheory/Monoidal/OfHasFiniteProducts.lean +++ b/Mathlib/CategoryTheory/Monoidal/OfHasFiniteProducts.lean @@ -73,7 +73,7 @@ open scoped MonoidalCategory @[ext] theorem tensor_ext {X Y Z : C} (f g : X ⟶ Y ⊗ Z) (w₁ : f ≫ prod.fst = g ≫ prod.fst) (w₂ : f ≫ prod.snd = g ≫ prod.snd) : f = g := - prod.hom_ext w₁ w₂ + Limits.prod.hom_ext w₁ w₂ @[simp] theorem tensorUnit : 𝟙_ C = ⊤_ C := rfl diff --git a/Mathlib/CategoryTheory/Products/Basic.lean b/Mathlib/CategoryTheory/Products/Basic.lean index 83d4a027d4ac0..dac0c3fe043ce 100644 --- a/Mathlib/CategoryTheory/Products/Basic.lean +++ b/Mathlib/CategoryTheory/Products/Basic.lean @@ -45,6 +45,11 @@ instance prod : Category.{max v₁ v₂} (C × D) where id X := ⟨𝟙 X.1, 𝟙 X.2⟩ comp f g := (f.1 ≫ g.1, f.2 ≫ g.2) +@[ext] +lemma prod.hom_ext {X Y : C × D} {f g : X ⟶ Y} (h₁ : f.1 = g.1) (h₂ : f.2 = g.2) : f = g := by + dsimp + ext <;> assumption + /-- Two rfl lemmas that cannot be generated by `@[simps]`. -/ @[simp] theorem prod_id (X : C) (Y : D) : 𝟙 (X, Y) = (𝟙 X, 𝟙 Y) := From 8e804ebb030600c8d597aeb073a0b902f63cde14 Mon Sep 17 00:00:00 2001 From: Matthew Robert Ballard Date: Thu, 17 Oct 2024 07:12:45 +0000 Subject: [PATCH 054/505] chore(ZMod.Basic): remove unnecessary `nolint` (#17840) This does not trigger the `unusedHavesSuffices` linter. --- Mathlib/Data/ZMod/Basic.lean | 1 - 1 file changed, 1 deletion(-) diff --git a/Mathlib/Data/ZMod/Basic.lean b/Mathlib/Data/ZMod/Basic.lean index d7eb347c8deee..b8c5302f15179 100644 --- a/Mathlib/Data/ZMod/Basic.lean +++ b/Mathlib/Data/ZMod/Basic.lean @@ -780,7 +780,6 @@ def inv : ∀ n : ℕ, ZMod n → ZMod n instance (n : ℕ) : Inv (ZMod n) := ⟨inv n⟩ -@[nolint unusedHavesSuffices] theorem inv_zero : ∀ n : ℕ, (0 : ZMod n)⁻¹ = 0 | 0 => Int.sign_zero | n + 1 => From 9caf52848a187af346984fdd790422b60183b41f Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Thu, 17 Oct 2024 07:12:46 +0000 Subject: [PATCH 055/505] chore: cleanup in Data/Prod/Basic (#17841) --- Mathlib/Data/List/NatAntidiagonal.lean | 2 +- Mathlib/Data/List/Sublists.lean | 2 +- Mathlib/Data/Multiset/Antidiagonal.lean | 2 +- Mathlib/Data/Prod/Basic.lean | 15 +++++---------- 4 files changed, 8 insertions(+), 13 deletions(-) diff --git a/Mathlib/Data/List/NatAntidiagonal.lean b/Mathlib/Data/List/NatAntidiagonal.lean index 3d11150f01d9e..b7ccdccf19f70 100644 --- a/Mathlib/Data/List/NatAntidiagonal.lean +++ b/Mathlib/Data/List/NatAntidiagonal.lean @@ -60,7 +60,7 @@ theorem nodup_antidiagonal (n : ℕ) : Nodup (antidiagonal n) := theorem antidiagonal_succ {n : ℕ} : antidiagonal (n + 1) = (0, n + 1) :: (antidiagonal n).map (Prod.map Nat.succ id) := by simp only [antidiagonal, range_succ_eq_map, map_cons, Nat.add_succ_sub_one, - Nat.add_zero, id, eq_self_iff_true, Nat.sub_zero, map_map, Prod.map_mk] + Nat.add_zero, id, eq_self_iff_true, Nat.sub_zero, map_map, Prod.map_apply] apply congr rfl (congr rfl _) ext; simp diff --git a/Mathlib/Data/List/Sublists.lean b/Mathlib/Data/List/Sublists.lean index 5521bdab8a956..bbe9b5bfedc30 100644 --- a/Mathlib/Data/List/Sublists.lean +++ b/Mathlib/Data/List/Sublists.lean @@ -378,7 +378,7 @@ theorem revzip_sublists (l : List α) : ∀ l₁ l₂, (l₁, l₂) ∈ revzip l · intro l₁ l₂ h rw [sublists_concat, reverse_append, zip_append (by simp), ← map_reverse, zip_map_right, zip_map_left] at * - simp only [Prod.mk.inj_iff, mem_map, mem_append, Prod.map_mk, Prod.exists] at h + simp only [Prod.mk.inj_iff, mem_map, mem_append, Prod.map_apply, Prod.exists] at h rcases h with (⟨l₁, l₂', h, rfl, rfl⟩ | ⟨l₁', l₂, h, rfl, rfl⟩) · rw [← append_assoc] exact (ih _ _ h).append_right _ diff --git a/Mathlib/Data/Multiset/Antidiagonal.lean b/Mathlib/Data/Multiset/Antidiagonal.lean index df02ceb0f2f41..dea72ad65efc7 100644 --- a/Mathlib/Data/Multiset/Antidiagonal.lean +++ b/Mathlib/Data/Multiset/Antidiagonal.lean @@ -82,7 +82,7 @@ theorem antidiagonal_eq_map_powerset [DecidableEq α] (s : Multiset α) : s.antidiagonal = s.powerset.map fun t ↦ (s - t, t) := by induction' s using Multiset.induction_on with a s hs · simp only [antidiagonal_zero, powerset_zero, Multiset.zero_sub, map_singleton] - · simp_rw [antidiagonal_cons, powerset_cons, map_add, hs, map_map, Function.comp, Prod.map_mk, + · simp_rw [antidiagonal_cons, powerset_cons, map_add, hs, map_map, Function.comp, Prod.map_apply, id, sub_cons, erase_cons_head] rw [add_comm] congr 1 diff --git a/Mathlib/Data/Prod/Basic.lean b/Mathlib/Data/Prod/Basic.lean index 125f987461805..2fbd64bdcf181 100644 --- a/Mathlib/Data/Prod/Basic.lean +++ b/Mathlib/Data/Prod/Basic.lean @@ -42,25 +42,20 @@ theorem snd_comp_mk (x : α) : Prod.snd ∘ (Prod.mk x : β → α × β) = id : theorem fst_comp_mk (x : α) : Prod.fst ∘ (Prod.mk x : β → α × β) = Function.const β x := rfl -@[simp, mfld_simps] -theorem map_mk (f : α → γ) (g : β → δ) (a : α) (b : β) : map f g (a, b) = (f a, g b) := - rfl +@[deprecated (since := "2024-10-17")] alias map_mk := map_apply + +attribute [mfld_simps] map_apply -- This was previously a `simp` lemma, but no longer is on the basis that it destructures the pair. -- See `map_apply`, `map_fst`, and `map_snd` for slightly weaker lemmas in the `simp` set. theorem map_apply' (f : α → γ) (g : β → δ) (p : α × β) : map f g p = (f p.1, g p.2) := rfl -#adaptation_note -/-- -After `nightly-2024-06-23`, the explicitness of `map_fst` and `map_snd` will be fixed and we can -change this back to `funext <| map_fst f g`. Also in `map_snd'` below. --/ theorem map_fst' (f : α → γ) (g : β → δ) : Prod.fst ∘ map f g = f ∘ Prod.fst := - funext <| @map_fst (f := f) (g := g) + funext <| map_fst f g theorem map_snd' (f : α → γ) (g : β → δ) : Prod.snd ∘ map f g = g ∘ Prod.snd := - funext <| @map_snd (f := f) (g := g) + funext <| map_snd f g /-- Composing a `Prod.map` with another `Prod.map` is equal to a single `Prod.map` of composed functions. From abde41d4f96c622fe866005b4b410f3cb3767183 Mon Sep 17 00:00:00 2001 From: mathlib4-update-dependencies-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Thu, 17 Oct 2024 07:52:57 +0000 Subject: [PATCH 056/505] chore: update Mathlib dependencies 2024-10-17 (#17856) This PR updates the Mathlib dependencies. --- lake-manifest.json | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lake-manifest.json b/lake-manifest.json index 968440a4a1bb7..1dc73cf0afe47 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,7 +5,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "d011c00010429728fab0bb3675b13f4bce63539e", + "rev": "cc0bc876eeef0518ddc1c8d3bd6f48cc83e68901", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", From 21bb8348b419e6365e8b174cebc3ce95714da087 Mon Sep 17 00:00:00 2001 From: Dagur Asgeirsson Date: Thu, 17 Oct 2024 08:31:01 +0000 Subject: [PATCH 057/505] feat(Condensed): characterisation of discrete (light) condensed sets and modules (#14027) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR provides a characterization of discrete condensed sets and modules, both in the light setting and the classical one. * Six equivalent conditions on a condensed set X to be discrete 1. The counit of the discrete-underlying adjunction applied to X is an isomorphism 2. X is in the essential image of the constant sheaf functor `Type (u+1) ⥤ CondensedSet.{u}` 3. X is in the essential image of the functor `Type (u+1) ⥤ CondensedSet.{u}` which takes a set to the sheaf of locally constant maps into it. 4. The counit of the locally-constant-underlying adjunction applied to X is an isomorphism. 5. X restricted to the coherent site of profinite sets is discrete as a sheaf 6. For every profinite set S written as a limit of finite sets, X maps S to the corresponding colimit. * A condensed module over a ring is discrete if and only if its underlying condensed set is, and therefore the analogues of the equivalent conditions above also characterize condensed modules as discrete * The analogues for the above for light condensed sets and modules --- Mathlib.lean | 1 + .../Category/ModuleCat/FilteredColimits.lean | 5 +- .../Condensed/Discrete/Characterization.lean | 264 ++++++++++++++++++ 3 files changed, 269 insertions(+), 1 deletion(-) create mode 100644 Mathlib/Condensed/Discrete/Characterization.lean diff --git a/Mathlib.lean b/Mathlib.lean index 14f4a5f1fa41b..39bb6b56e139b 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2078,6 +2078,7 @@ import Mathlib.Computability.TuringMachine import Mathlib.Condensed.Basic import Mathlib.Condensed.CartesianClosed import Mathlib.Condensed.Discrete.Basic +import Mathlib.Condensed.Discrete.Characterization import Mathlib.Condensed.Discrete.Colimit import Mathlib.Condensed.Discrete.LocallyConstant import Mathlib.Condensed.Discrete.Module diff --git a/Mathlib/Algebra/Category/ModuleCat/FilteredColimits.lean b/Mathlib/Algebra/Category/ModuleCat/FilteredColimits.lean index 16abcd3247283..58c0fe8253d3a 100644 --- a/Mathlib/Algebra/Category/ModuleCat/FilteredColimits.lean +++ b/Mathlib/Algebra/Category/ModuleCat/FilteredColimits.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Justus Springer -/ import Mathlib.Algebra.Category.Grp.FilteredColimits -import Mathlib.Algebra.Category.ModuleCat.Basic +import Mathlib.Algebra.Category.ModuleCat.Colimits /-! # The forgetful functor from `R`-modules preserves filtered colimits. @@ -184,6 +184,9 @@ instance forgetPreservesFilteredColimits : PreservesFilteredColimits (forget (Mo Limits.compPreservesFilteredColimits (forget₂ (ModuleCat R) AddCommGrp) (forget AddCommGrp) +instance forgetReflectsFilteredColimits : ReflectsFilteredColimits (forget (ModuleCat.{u} R)) where + reflects_filtered_colimits _ := { reflectsColimit := reflectsColimitOfReflectsIsomorphisms _ _ } + end end ModuleCat.FilteredColimits diff --git a/Mathlib/Condensed/Discrete/Characterization.lean b/Mathlib/Condensed/Discrete/Characterization.lean new file mode 100644 index 0000000000000..04c5cbbfc2c17 --- /dev/null +++ b/Mathlib/Condensed/Discrete/Characterization.lean @@ -0,0 +1,264 @@ +/- +Copyright (c) 2024 Dagur Asgeirsson. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Dagur Asgeirsson +-/ +import Mathlib.Condensed.Discrete.Colimit +import Mathlib.Condensed.Discrete.Module +/-! + +# Characterizing discrete condensed sets and `R`-modules. + +This file proves a characterization of discrete condensed sets, discrete condensed `R`-modules over +a ring `R`, discrete light condensed sets, and discrete light condensed `R`-modules over a ring `R`. +see `CondensedSet.isDiscrete_tfae`, `CondensedMod.isDiscrete_tfae`, `LightCondSet.isDiscrete_tfae`, +and `LightCondMod.isDiscrete_tfae`. + +Informally, we can say: The following conditions characterize a condensed set `X` as discrete +(`CondensedSet.isDiscrete_tfae`): + +1. There exists a set `X'` and an isomorphism `X ≅ cst X'`, where `cst X'` denotes the constant + sheaf on `X'`. +2. The counit induces an isomorphism `cst X(*) ⟶ X`. +3. There exists a set `X'` and an isomorphism `X ≅ LocallyConstant · X'`. +4. The counit induces an isomorphism `LocallyConstant · X(*) ⟶ X`. +5. For every profinite set `S = limᵢSᵢ`, the canonical map `colimᵢX(Sᵢ) ⟶ X(S)` is an isomorphism. + +The analogues for light condensed sets, condensed `R`-modules over any ring, and light +condensed `R`-modules are nearly identical (`CondensedMod.isDiscrete_tfae`, +`LightCondSet.isDiscrete_tfae`, and `LightCondMod.isDiscrete_tfae`). +-/ + +universe u + +open CategoryTheory Limits Functor FintypeCat + +attribute [local instance] ConcreteCategory.instFunLike + +namespace Condensed + +variable {C : Type*} [Category C] [HasWeakSheafify (coherentTopology CompHaus.{u}) C] + +/-- +A condensed object is *discrete* if it is constant as a sheaf, i.e. isomorphic to a constant sheaf. +-/ +abbrev IsDiscrete (X : Condensed.{u} C) := X.IsConstant (coherentTopology CompHaus) + +end Condensed + +namespace CondensedSet + +open CompHausLike.LocallyConstant + +lemma mem_locallyContant_essImage_of_isColimit_mapCocone (X : CondensedSet.{u}) + (h : ∀ S : Profinite.{u}, IsColimit <| + (profiniteToCompHaus.op ⋙ X.val).mapCocone S.asLimitCone.op) : + X ∈ CondensedSet.LocallyConstant.functor.essImage := by + let e : CondensedSet.{u} ≌ Sheaf (coherentTopology Profinite) _ := + (Condensed.ProfiniteCompHaus.equivalence (Type (u + 1))).symm + let i : (e.functor.obj X).val ≅ (e.functor.obj (LocallyConstant.functor.obj _)).val := + Condensed.isoLocallyConstantOfIsColimit _ h + exact ⟨_, ⟨e.functor.preimageIso ((sheafToPresheaf _ _).preimageIso i.symm)⟩⟩ + +/-- +`CondensedSet.LocallyConstant.functor` is left adjoint to the forgetful functor from condensed +sets to sets. +-/ +noncomputable abbrev LocallyConstant.adjunction : + CondensedSet.LocallyConstant.functor ⊣ Condensed.underlying (Type (u+1)) := + CompHausLike.LocallyConstant.adjunction _ _ + +open Condensed + +open CondensedSet.LocallyConstant List in +theorem isDiscrete_tfae (X : CondensedSet.{u}) : + TFAE + [ X.IsDiscrete + , IsIso ((Condensed.discreteUnderlyingAdj _).counit.app X) + , X ∈ (Condensed.discrete _).essImage + , X ∈ CondensedSet.LocallyConstant.functor.essImage + , IsIso (CondensedSet.LocallyConstant.adjunction.counit.app X) + , Sheaf.IsConstant (coherentTopology Profinite) + ((Condensed.ProfiniteCompHaus.equivalence _).inverse.obj X) + , ∀ S : Profinite.{u}, Nonempty + (IsColimit <| (profiniteToCompHaus.op ⋙ X.val).mapCocone S.asLimitCone.op) + ] := by + tfae_have 1 ↔ 2 := Sheaf.isConstant_iff_isIso_counit_app _ _ _ + tfae_have 1 ↔ 3 := ⟨fun ⟨h⟩ ↦ h, fun h ↦ ⟨h⟩⟩ + tfae_have 1 ↔ 4 := Sheaf.isConstant_iff_mem_essImage _ CompHaus.isTerminalPUnit adjunction _ + tfae_have 1 ↔ 5 := + have : functor.Faithful := inferInstance + have : functor.Full := inferInstance + -- These `have` statements above shouldn't be needed, but they are. + Sheaf.isConstant_iff_isIso_counit_app' _ CompHaus.isTerminalPUnit adjunction _ + tfae_have 1 ↔ 6 := + (Sheaf.isConstant_iff_of_equivalence (coherentTopology Profinite) + (coherentTopology CompHaus) profiniteToCompHaus Profinite.isTerminalPUnit + CompHaus.isTerminalPUnit _).symm + tfae_have 7 → 4 := fun h ↦ + mem_locallyContant_essImage_of_isColimit_mapCocone X (fun S ↦ (h S).some) + tfae_have 4 → 7 := fun ⟨Y, ⟨i⟩⟩ S ↦ + ⟨IsColimit.mapCoconeEquiv (isoWhiskerLeft profiniteToCompHaus.op + ((sheafToPresheaf _ _).mapIso i)) + (Condensed.isColimitLocallyConstantPresheafDiagram Y S)⟩ + tfae_finish + +end CondensedSet + +namespace CondensedMod + +variable (R : Type (u+1)) [Ring R] + +lemma isDiscrete_iff_isDiscrete_forget (M : CondensedMod R) : + M.IsDiscrete ↔ ((Condensed.forget R).obj M).IsDiscrete := + Sheaf.isConstant_iff_forget (coherentTopology CompHaus) + (forget (ModuleCat R)) M CompHaus.isTerminalPUnit + +instance : HasLimitsOfSize.{u, u+1} (ModuleCat.{u+1} R) := + hasLimitsOfSizeShrink.{u, u+1, u+1, u+1} _ + +open CondensedMod.LocallyConstant List in +theorem isDiscrete_tfae (M : CondensedMod.{u} R) : + TFAE + [ M.IsDiscrete + , IsIso ((Condensed.discreteUnderlyingAdj _).counit.app M) + , M ∈ (Condensed.discrete _).essImage + , M ∈ (CondensedMod.LocallyConstant.functor R).essImage + , IsIso ((CondensedMod.LocallyConstant.adjunction R).counit.app M) + , Sheaf.IsConstant (coherentTopology Profinite) + ((Condensed.ProfiniteCompHaus.equivalence _).inverse.obj M) + , ∀ S : Profinite.{u}, Nonempty + (IsColimit <| (profiniteToCompHaus.op ⋙ M.val).mapCocone S.asLimitCone.op) + ] := by + tfae_have 1 ↔ 2 := Sheaf.isConstant_iff_isIso_counit_app _ _ _ + tfae_have 1 ↔ 3 := ⟨fun ⟨h⟩ ↦ h, fun h ↦ ⟨h⟩⟩ + tfae_have 1 ↔ 4 := Sheaf.isConstant_iff_mem_essImage _ CompHaus.isTerminalPUnit (adjunction R) _ + tfae_have 1 ↔ 5 := + have : (functor R).Faithful := inferInstance + have : (functor R).Full := inferInstance + -- These `have` statements above shouldn't be needed, but they are. + Sheaf.isConstant_iff_isIso_counit_app' _ CompHaus.isTerminalPUnit (adjunction R) _ + tfae_have 1 ↔ 6 := + (Sheaf.isConstant_iff_of_equivalence (coherentTopology Profinite) + (coherentTopology CompHaus) profiniteToCompHaus Profinite.isTerminalPUnit + CompHaus.isTerminalPUnit _).symm + tfae_have 7 → 1 := by + intro h + rw [isDiscrete_iff_isDiscrete_forget, ((CondensedSet.isDiscrete_tfae _).out 0 6:)] + intro S + letI : PreservesFilteredColimitsOfSize.{u, u} (forget (ModuleCat R)) := + preservesFilteredColimitsOfSizeShrink.{u, u+1, u, u+1} _ + exact ⟨isColimitOfPreserves (forget (ModuleCat R)) (h S).some⟩ + tfae_have 1 → 7 := by + intro h S + rw [isDiscrete_iff_isDiscrete_forget, ((CondensedSet.isDiscrete_tfae _).out 0 6:)] at h + letI : ReflectsFilteredColimitsOfSize.{u, u} (forget (ModuleCat R)) := + reflectsFilteredColimitsOfSizeShrink.{u, u+1, u, u+1} _ + exact ⟨isColimitOfReflects (forget (ModuleCat R)) (h S).some⟩ + tfae_finish + +end CondensedMod + +namespace LightCondensed + +variable {C : Type*} [Category C] [HasWeakSheafify (coherentTopology LightProfinite.{u}) C] + +/-- +A light condensed object is *discrete* if it is constant as a sheaf, i.e. isomorphic to a constant +sheaf. +-/ +abbrev IsDiscrete (X : LightCondensed.{u} C) := X.IsConstant (coherentTopology LightProfinite) + +end LightCondensed + +namespace LightCondSet + +lemma mem_locallyContant_essImage_of_isColimit_mapCocone (X : LightCondSet.{u}) + (h : ∀ S : LightProfinite.{u}, IsColimit <| + X.val.mapCocone (coconeRightOpOfCone S.asLimitCone)) : + X ∈ LightCondSet.LocallyConstant.functor.essImage := by + let i : X.val ≅ (LightCondSet.LocallyConstant.functor.obj _).val := + LightCondensed.isoLocallyConstantOfIsColimit _ h + exact ⟨_, ⟨((sheafToPresheaf _ _).preimageIso i.symm)⟩⟩ + +/-- +`LightCondSet.LocallyConstant.functor` is left adjoint to the forgetful functor from light condensed +sets to sets. +-/ +noncomputable abbrev LocallyConstant.adjunction : + LightCondSet.LocallyConstant.functor ⊣ LightCondensed.underlying (Type u) := + CompHausLike.LocallyConstant.adjunction _ _ + +open LightCondSet.LocallyConstant List in +theorem isDiscrete_tfae (X : LightCondSet.{u}) : + TFAE + [ X.IsDiscrete + , IsIso ((LightCondensed.discreteUnderlyingAdj _).counit.app X) + , X ∈ (LightCondensed.discrete _).essImage + , X ∈ LightCondSet.LocallyConstant.functor.essImage + , IsIso (LightCondSet.LocallyConstant.adjunction.counit.app X) + , ∀ S : LightProfinite.{u}, Nonempty + (IsColimit <| X.val.mapCocone (coconeRightOpOfCone S.asLimitCone)) + ] := by + tfae_have 1 ↔ 2 := Sheaf.isConstant_iff_isIso_counit_app _ _ _ + tfae_have 1 ↔ 3 := ⟨fun ⟨h⟩ ↦ h, fun h ↦ ⟨h⟩⟩ + tfae_have 1 ↔ 4 := Sheaf.isConstant_iff_mem_essImage _ LightProfinite.isTerminalPUnit adjunction X + tfae_have 1 ↔ 5 := + have : functor.Faithful := inferInstance + have : functor.Full := inferInstance + -- These `have` statements above shouldn't be needed, but they are. + Sheaf.isConstant_iff_isIso_counit_app' _ LightProfinite.isTerminalPUnit adjunction X + tfae_have 6 → 4 := fun h ↦ + mem_locallyContant_essImage_of_isColimit_mapCocone X (fun S ↦ (h S).some) + tfae_have 4 → 6 := fun ⟨Y, ⟨i⟩⟩ S ↦ + ⟨IsColimit.mapCoconeEquiv ((sheafToPresheaf _ _).mapIso i) + (LightCondensed.isColimitLocallyConstantPresheafDiagram Y S)⟩ + tfae_finish + +end LightCondSet + +namespace LightCondMod + +variable (R : Type u) [Ring R] + +lemma isDiscrete_iff_isDiscrete_forget (M : LightCondMod R) : + M.IsDiscrete ↔ ((LightCondensed.forget R).obj M).IsDiscrete := + Sheaf.isConstant_iff_forget (coherentTopology LightProfinite) + (forget (ModuleCat R)) M LightProfinite.isTerminalPUnit + +open LightCondMod.LocallyConstant List in +theorem isDiscrete_tfae (M : LightCondMod.{u} R) : + TFAE + [ M.IsDiscrete + , IsIso ((LightCondensed.discreteUnderlyingAdj _).counit.app M) + , M ∈ (LightCondensed.discrete _).essImage + , M ∈ (LightCondMod.LocallyConstant.functor R).essImage + , IsIso ((LightCondMod.LocallyConstant.adjunction R).counit.app M) + , ∀ S : LightProfinite.{u}, Nonempty + (IsColimit <| M.val.mapCocone (coconeRightOpOfCone S.asLimitCone)) + ] := by + tfae_have 1 ↔ 2 := Sheaf.isConstant_iff_isIso_counit_app _ _ _ + tfae_have 1 ↔ 3 := ⟨fun ⟨h⟩ ↦ h, fun h ↦ ⟨h⟩⟩ + tfae_have 1 ↔ 4 := Sheaf.isConstant_iff_mem_essImage _ + LightProfinite.isTerminalPUnit (adjunction R) _ + tfae_have 1 ↔ 5 := + have : (functor R).Faithful := inferInstance + have : (functor R).Full := inferInstance + -- These `have` statements above shouldn't be needed, but they are. + Sheaf.isConstant_iff_isIso_counit_app' _ LightProfinite.isTerminalPUnit (adjunction R) _ + tfae_have 6 → 1 := by + intro h + rw [isDiscrete_iff_isDiscrete_forget, ((LightCondSet.isDiscrete_tfae _).out 0 5:)] + intro S + letI : PreservesFilteredColimitsOfSize.{0, 0} (forget (ModuleCat R)) := + preservesFilteredColimitsOfSizeShrink.{0, u, 0, u} _ + exact ⟨isColimitOfPreserves (forget (ModuleCat R)) (h S).some⟩ + tfae_have 1 → 6 := by + intro h S + rw [isDiscrete_iff_isDiscrete_forget, ((LightCondSet.isDiscrete_tfae _).out 0 5:)] at h + letI : ReflectsFilteredColimitsOfSize.{0, 0} (forget (ModuleCat R)) := + reflectsFilteredColimitsOfSizeShrink.{0, u, 0, u} _ + exact ⟨isColimitOfReflects (forget (ModuleCat R)) (h S).some⟩ + tfae_finish + +end LightCondMod From 9dba3dd644fa6310b6d1f0a7143491803d5005a4 Mon Sep 17 00:00:00 2001 From: EdwardWatine <02ewatine@gmail.com> Date: Thu, 17 Oct 2024 08:31:03 +0000 Subject: [PATCH 058/505] feat: define the ordinary hypergeometric function (#17455) Add the definition of the ordinary hypergeometric series, and show that its radius is one. Part of this requires an additional theorem about ascPochhammer, which is found in the Pochhammer file. Closely follows the exponential series definition. #15966 Co-authored-by: Alvan Caleb Arulandu --- Mathlib.lean | 2 + Mathlib/Analysis/Analytic/OfScalars.lean | 131 +++++++++++ .../OrdinaryHypergeometric.lean | 209 ++++++++++++++++++ Mathlib/Analysis/SpecificLimits/RCLike.lean | 15 ++ Mathlib/RingTheory/Polynomial/Pochhammer.lean | 40 +++- 5 files changed, 394 insertions(+), 3 deletions(-) create mode 100644 Mathlib/Analysis/Analytic/OfScalars.lean create mode 100644 Mathlib/Analysis/SpecialFunctions/OrdinaryHypergeometric.lean diff --git a/Mathlib.lean b/Mathlib.lean index 39bb6b56e139b..1170098fd659e 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -942,6 +942,7 @@ import Mathlib.Analysis.Analytic.Inverse import Mathlib.Analysis.Analytic.IsolatedZeros import Mathlib.Analysis.Analytic.Linear import Mathlib.Analysis.Analytic.Meromorphic +import Mathlib.Analysis.Analytic.OfScalars import Mathlib.Analysis.Analytic.Polynomial import Mathlib.Analysis.Analytic.RadiusLiminf import Mathlib.Analysis.Analytic.Uniqueness @@ -1363,6 +1364,7 @@ import Mathlib.Analysis.SpecialFunctions.Log.ERealExp import Mathlib.Analysis.SpecialFunctions.Log.Monotone import Mathlib.Analysis.SpecialFunctions.Log.NegMulLog import Mathlib.Analysis.SpecialFunctions.NonIntegrable +import Mathlib.Analysis.SpecialFunctions.OrdinaryHypergeometric import Mathlib.Analysis.SpecialFunctions.PolarCoord import Mathlib.Analysis.SpecialFunctions.PolynomialExp import Mathlib.Analysis.SpecialFunctions.Polynomials diff --git a/Mathlib/Analysis/Analytic/OfScalars.lean b/Mathlib/Analysis/Analytic/OfScalars.lean new file mode 100644 index 0000000000000..531079814dcc6 --- /dev/null +++ b/Mathlib/Analysis/Analytic/OfScalars.lean @@ -0,0 +1,131 @@ +/- +Copyright (c) 2024 Edward Watine. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Edward Watine +-/ +import Mathlib.Analysis.Analytic.Basic + + +/-! +# Scalar series + +This file contains API for analytic functions `∑ cᵢ • xⁱ` defined in terms of scalars +`c₀, c₁, c₂, …`. +## Main definitions / results: + * `FormalMultilinearSeries.ofScalars`: the formal power series `∑ cᵢ • xⁱ`. + * `FormalMultilinearSeries.ofScalars_radius_eq_of_tendsto`: the ratio test for an analytic function + defined in terms of a formal power series `∑ cᵢ • xⁱ`. +-/ + +namespace FormalMultilinearSeries + +section Field + +variable {𝕜 : Type*} (E : Type*) [Field 𝕜] [Ring E] [Algebra 𝕜 E] [TopologicalSpace E] + [TopologicalRing E] (c : ℕ → 𝕜) + +/-- Formal power series of `∑ cᵢ • xⁱ` for some scalar field `𝕜` and ring algebra `E`-/ +def ofScalars : FormalMultilinearSeries 𝕜 E E := + fun n ↦ c n • ContinuousMultilinearMap.mkPiAlgebraFin 𝕜 n E + +variable {E} + +/-- The sum of the formal power series. Takes the value `0` outside the radius of convergence. -/ +noncomputable def ofScalarsSum (x : E) := (ofScalars E c).sum x + +theorem ofScalars_apply_eq (x : E) (n : ℕ) : + ofScalars E c n (fun _ ↦ x) = c n • x ^ n := by + simp [ofScalars] + +/-- This naming follows the convention of `NormedSpace.expSeries_apply_eq'`. -/ +theorem ofScalars_apply_eq' (x : E) : + (fun n ↦ ofScalars E c n (fun _ ↦ x)) = fun n ↦ c n • x ^ n := by + simp [ofScalars] + +theorem ofScalars_sum_eq (x : E) : ofScalarsSum c x = + ∑' n, c n • x ^ n := tsum_congr fun n => ofScalars_apply_eq c x n + +theorem ofScalarsSum_eq_tsum : ofScalarsSum c = + fun (x : E) => ∑' n : ℕ, c n • x ^ n := funext (ofScalars_sum_eq c) + +@[simp] +theorem ofScalars_op [T2Space E] (x : E) : + ofScalarsSum c (MulOpposite.op x) = MulOpposite.op (ofScalarsSum c x) := by + simp [ofScalars, ofScalars_sum_eq, ← MulOpposite.op_pow, ← MulOpposite.op_smul, tsum_op] + +@[simp] +theorem ofScalars_unop [T2Space E] (x : Eᵐᵒᵖ) : + ofScalarsSum c (MulOpposite.unop x) = MulOpposite.unop (ofScalarsSum c x) := by + simp [ofScalars, ofScalars_sum_eq, ← MulOpposite.unop_pow, ← MulOpposite.unop_smul, tsum_unop] + +@[simp] +theorem ofScalars_eq_zero [Nontrivial E] (n : ℕ) : ofScalars E c n = 0 ↔ c n = 0 := by + rw [ofScalars, smul_eq_zero (c := c n) (x := ContinuousMultilinearMap.mkPiAlgebraFin 𝕜 n E)] + refine or_iff_left (ContinuousMultilinearMap.ext_iff.1.mt <| not_forall_of_exists_not ?_) + use fun _ ↦ 1 + simp + +end Field + +section Normed + +open Filter +open scoped Topology + +variable {𝕜 : Type*} (E : Type*) [NontriviallyNormedField 𝕜] [NormedRing E] + [NormedAlgebra 𝕜 E] (c : ℕ → 𝕜) (n : ℕ) + +theorem ofScalars_norm_eq_mul : + ‖ofScalars E c n‖ = ‖c n‖ * ‖ContinuousMultilinearMap.mkPiAlgebraFin 𝕜 n E‖ := by + rw [ofScalars, norm_smul (c n) (ContinuousMultilinearMap.mkPiAlgebraFin 𝕜 n E)] + +theorem ofScalars_norm_le (hn : n > 0) : ‖ofScalars E c n‖ ≤ ‖c n‖ := by + simp only [ofScalars_norm_eq_mul] + exact (mul_le_of_le_one_right (norm_nonneg _) + (ContinuousMultilinearMap.norm_mkPiAlgebraFin_le_of_pos hn)) + +@[simp] +theorem ofScalars_norm [NormOneClass E] : ‖ofScalars E c n‖ = ‖c n‖ := by + simp [ofScalars_norm_eq_mul] + +/-- The radius of convergence of a scalar series is the inverse of the ratio of the norms of the +coefficients. -/ +theorem ofScalars_radius_eq_inv_of_tendsto [NormOneClass E] {r : NNReal} (hr : r ≠ 0) + (hc : Tendsto (fun n ↦ ‖c n.succ‖ / ‖c n‖) atTop (𝓝 r)) : + (ofScalars E c).radius = ENNReal.ofNNReal r⁻¹ := by + have hc' {r' : NNReal} (hr' : (r' : ℝ) ≠ 0) : + Tendsto (fun n ↦ ‖‖ofScalars E c (n + 1)‖ * r' ^ (n + 1)‖ / + ‖‖ofScalars E c n‖ * r' ^ n‖) atTop (𝓝 ↑(r' * r)) := by + simp_rw [norm_mul, norm_norm, ofScalars_norm, mul_div_mul_comm, ← norm_div, pow_succ, + mul_div_right_comm, div_self (pow_ne_zero _ hr'), one_mul, norm_div, NNReal.norm_eq] + exact mul_comm r' r ▸ Filter.Tendsto.mul hc tendsto_const_nhds + apply le_antisymm <;> refine ENNReal.le_of_forall_nnreal_lt (fun r' hr' ↦ ?_) + · rw [ENNReal.coe_le_coe, NNReal.le_inv_iff_mul_le hr] + have := FormalMultilinearSeries.summable_norm_mul_pow _ hr' + contrapose! this + have hrz : (r' : ℝ) ≠ 0 := by aesop + apply not_summable_of_ratio_test_tendsto_gt_one this + exact hc' (by aesop) + · rw [ENNReal.coe_lt_coe, NNReal.lt_inv_iff_mul_lt hr] at hr' + by_cases hrz : r' = 0 + · simp [hrz] + · apply FormalMultilinearSeries.le_radius_of_summable_norm + refine summable_of_ratio_test_tendsto_lt_one hr' ?_ <| hc' (NNReal.coe_ne_zero.2 hrz) + refine (Tendsto.eventually_ne hc (NNReal.coe_ne_zero.2 hr)).mp (Eventually.of_forall ?_) + simp_rw [div_ne_zero_iff, ofScalars_norm, mul_ne_zero_iff] + aesop + +/-- A convenience lemma restating the result of `ofScalars_radius_eq_inv_of_tendsto` under +the inverse ratio. -/ +theorem ofScalars_radius_eq_of_tendsto [NormOneClass E] {r : NNReal} (hr : r ≠ 0) + (hc : Tendsto (fun n ↦ ‖c n‖ / ‖c n.succ‖) atTop (𝓝 r)) : + (ofScalars E c).radius = ENNReal.ofNNReal r := by + suffices Tendsto (fun n ↦ ‖c n.succ‖ / ‖c n‖) atTop (𝓝 r⁻¹) by + convert ofScalars_radius_eq_inv_of_tendsto E c (inv_ne_zero hr) this + simp + convert (continuousAt_inv₀ <| NNReal.coe_ne_zero.mpr hr).tendsto.comp hc + simp + +end Normed + +end FormalMultilinearSeries diff --git a/Mathlib/Analysis/SpecialFunctions/OrdinaryHypergeometric.lean b/Mathlib/Analysis/SpecialFunctions/OrdinaryHypergeometric.lean new file mode 100644 index 0000000000000..bbbd460e9abe3 --- /dev/null +++ b/Mathlib/Analysis/SpecialFunctions/OrdinaryHypergeometric.lean @@ -0,0 +1,209 @@ +/- +Copyright (c) 2024 Edward Watine. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Edward Watine +-/ +import Mathlib.Analysis.Analytic.OfScalars +import Mathlib.Analysis.SpecificLimits.RCLike + +/-! +# Ordinary hypergeometric function in a Banach algebra + +In this file, we define `ordinaryHypergeometric`, the _ordinary_ or _Gaussian_ hypergeometric +function in a topological algebra `𝔸` over a field `𝕂` given by: $$ +_2\mathrm{F}_1(a\ b\ c : \mathbb{K}, x : \mathbb{A}) = \sum_{n=0}^{\infty}\frac{(a)_n(b)_n}{(c)_n} +\frac{x^n}{n!} \,, +$$ +with $(a)_n$ is the ascending Pochhammer symbol (see `ascPochhammer`). + +This file contains the basic definitions over a general field `𝕂` and notation for `₂F₁`, +as well as showing that terms of the series are zero if any of the `(a b c : 𝕂)` are sufficiently +large non-positive integers, rendering the series finite. In this file "sufficiently large" means +that `-n < a` for the `n`-th term, and similarly for `b` and `c`. + +- `ordinaryHypergeometricSeries` is the `FormalMultilinearSeries` given above for some `(a b c : 𝕂)` +- `ordinaryHypergeometric` is the sum of the series for some `(x : 𝔸)` +- `ordinaryHypergeometricSeries_eq_zero_of_nonpos_int` shows that the `n`-th term of the series is +zero if any of the parameters are sufficiently large non-positive integers + +## `[RCLike 𝕂]` + +If we have `[RCLike 𝕂]`, then we show that the latter result is an iff, and hence prove that the +radius of convergence of the series is unity if the series is infinite, or `⊤` otherwise. + +- `ordinaryHypergeometricSeries_eq_zero_iff` is iff variant of +`ordinaryHypergeometricSeries_eq_zero_of_nonpos_int` +- `ordinaryHypergeometricSeries_radius_eq_one` proves that the radius of convergence of the +`ordinaryHypergeometricSeries` is unity under non-trivial parameters + +## Notation + +`₂F₁` is notation for `ordinaryHypergeometric`. + +## References + +See . + +## Tags + +hypergeometric, gaussian, ordinary +-/ + +open Nat FormalMultilinearSeries + +section Field + +variable {𝕂 : Type*} (𝔸 : Type*) [Field 𝕂] [Ring 𝔸] [Algebra 𝕂 𝔸] [TopologicalSpace 𝔸] + [TopologicalRing 𝔸] + +/-- The coefficients in the ordinary hypergeometric sum. -/ +noncomputable abbrev ordinaryHypergeometricCoefficient (a b c : 𝕂) (n : ℕ) := ((n !⁻¹ : 𝕂) * + (ascPochhammer 𝕂 n).eval a * (ascPochhammer 𝕂 n).eval b * ((ascPochhammer 𝕂 n).eval c)⁻¹) + +/-- `ordinaryHypergeometricSeries 𝔸 (a b c : 𝕂)` is a `FormalMultilinearSeries`. +Its sum is the `ordinaryHypergeometric` map. -/ +noncomputable def ordinaryHypergeometricSeries (a b c : 𝕂) : FormalMultilinearSeries 𝕂 𝔸 𝔸 := + ofScalars 𝔸 (ordinaryHypergeometricCoefficient a b c) + +variable {𝔸} (a b c : 𝕂) + +/-- `ordinaryHypergeometric (a b c : 𝕂) : 𝔸 → 𝔸` is the ordinary hypergeometric map, defined as the +sum of the `FormalMultilinearSeries` `ordinaryHypergeometricSeries 𝔸 a b c`. + +Note that this takes the junk value `0` outside the radius of convergence. +-/ +noncomputable def ordinaryHypergeometric (x : 𝔸) : 𝔸 := + (ordinaryHypergeometricSeries 𝔸 a b c).sum x + +@[inherit_doc] +notation "₂F₁" => ordinaryHypergeometric + +theorem ordinaryHypergeometricSeries_apply_eq (x : 𝔸) (n : ℕ) : + (ordinaryHypergeometricSeries 𝔸 a b c n fun _ => x) = + ((n !⁻¹ : 𝕂) * (ascPochhammer 𝕂 n).eval a * (ascPochhammer 𝕂 n).eval b * + ((ascPochhammer 𝕂 n).eval c)⁻¹ ) • x ^ n := by + rw [ordinaryHypergeometricSeries, ofScalars_apply_eq] + +/-- This naming follows the convention of `NormedSpace.expSeries_apply_eq'`. -/ +theorem ordinaryHypergeometricSeries_apply_eq' (x : 𝔸) : + (fun n => ordinaryHypergeometricSeries 𝔸 a b c n fun _ => x) = + fun n => ((n !⁻¹ : 𝕂) * (ascPochhammer 𝕂 n).eval a * (ascPochhammer 𝕂 n).eval b * + ((ascPochhammer 𝕂 n).eval c)⁻¹ ) • x ^ n := by + rw [ordinaryHypergeometricSeries, ofScalars_apply_eq'] + +theorem ordinaryHypergeometric_sum_eq (x : 𝔸) : (ordinaryHypergeometricSeries 𝔸 a b c).sum x = + ∑' n : ℕ, ((n !⁻¹ : 𝕂) * (ascPochhammer 𝕂 n).eval a * (ascPochhammer 𝕂 n).eval b * + ((ascPochhammer 𝕂 n).eval c)⁻¹ ) • x ^ n := + tsum_congr fun n => ordinaryHypergeometricSeries_apply_eq a b c x n + +theorem ordinaryHypergeometric_eq_tsum : ₂F₁ a b c = + fun (x : 𝔸) => ∑' n : ℕ, ((n !⁻¹ : 𝕂) * (ascPochhammer 𝕂 n).eval a * + (ascPochhammer 𝕂 n).eval b * ((ascPochhammer 𝕂 n).eval c)⁻¹ ) • x ^ n := + funext (ordinaryHypergeometric_sum_eq a b c) + +theorem ordinaryHypergeometricSeries_apply_zero (n : ℕ) : + (ordinaryHypergeometricSeries 𝔸 a b c n fun _ => 0) = Pi.single (f := fun _ => 𝔸) 0 1 n := by + rw [ordinaryHypergeometricSeries, ofScalars_apply_eq, ordinaryHypergeometricCoefficient] + cases n <;> simp + +@[simp] +theorem ordinaryHypergeometric_zero : ₂F₁ a b c (0 : 𝔸) = 1 := by + simp [ordinaryHypergeometric_eq_tsum, ← ordinaryHypergeometricSeries_apply_eq, + ordinaryHypergeometricSeries_apply_zero] + +theorem ordinaryHypergeometricSeries_symm : + ordinaryHypergeometricSeries 𝔸 a b c = ordinaryHypergeometricSeries 𝔸 b a c := by + unfold ordinaryHypergeometricSeries ordinaryHypergeometricCoefficient + simp [mul_assoc, mul_left_comm] + +/-- If any parameter to the series is a sufficiently large nonpositive integer, then the series +term is zero. -/ +lemma ordinaryHypergeometricSeries_eq_zero_of_neg_nat {n k : ℕ} (habc : k = -a ∨ k = -b ∨ k = -c) + (hk : k < n) : ordinaryHypergeometricSeries 𝔸 a b c n = 0 := by + rw [ordinaryHypergeometricSeries, ofScalars] + rcases habc with h | h | h + all_goals + ext + simp [(ascPochhammer_eval_eq_zero_iff n _).2 ⟨k, hk, h⟩] + +end Field + +section RCLike + +open Asymptotics Filter Real Set Nat + +open scoped Topology + +variable {𝕂 : Type*} (𝔸 : Type*) [RCLike 𝕂] [NormedDivisionRing 𝔸] [NormedAlgebra 𝕂 𝔸] + (a b c : 𝕂) + +theorem ordinaryHypergeometric_radius_top_of_neg_nat₁ {k : ℕ} : + (ordinaryHypergeometricSeries 𝔸 (-(k : 𝕂)) b c).radius = ⊤ := by + refine FormalMultilinearSeries.radius_eq_top_of_forall_image_add_eq_zero _ (1 + k) fun n ↦ ?_ + exact ordinaryHypergeometricSeries_eq_zero_of_neg_nat (-(k : 𝕂)) b c (by aesop) (by omega) + +theorem ordinaryHypergeometric_radius_top_of_neg_nat₂ {k : ℕ} : + (ordinaryHypergeometricSeries 𝔸 a (-(k : 𝕂)) c).radius = ⊤ := by + rw [ordinaryHypergeometricSeries_symm] + exact ordinaryHypergeometric_radius_top_of_neg_nat₁ 𝔸 a c + +theorem ordinaryHypergeometric_radius_top_of_neg_nat₃ {k : ℕ} : + (ordinaryHypergeometricSeries 𝔸 a b (-(k : 𝕂))).radius = ⊤ := by + refine FormalMultilinearSeries.radius_eq_top_of_forall_image_add_eq_zero _ (1 + k) fun n ↦ ?_ + exact ordinaryHypergeometricSeries_eq_zero_of_neg_nat a b (-(k : 𝕂)) (by aesop) (by omega) + +/-- An iff variation on `ordinaryHypergeometricSeries_eq_zero_of_nonpos_int` for `[RCLike 𝕂]`. -/ +lemma ordinaryHypergeometricSeries_eq_zero_iff (n : ℕ) : + ordinaryHypergeometricSeries 𝔸 a b c n = 0 ↔ ∃ k < n, k = -a ∨ k = -b ∨ k = -c := by + refine ⟨fun h ↦ ?_, fun zero ↦ ?_⟩ + · rw [ordinaryHypergeometricSeries, ofScalars_eq_zero] at h + simp only [_root_.mul_eq_zero, inv_eq_zero] at h + rcases h with ((hn | h) | h) | h + · simp [Nat.factorial_ne_zero] at hn + all_goals + obtain ⟨kn, hkn, hn⟩ := (ascPochhammer_eval_eq_zero_iff _ _).1 h + exact ⟨kn, hkn, by tauto⟩ + · obtain ⟨_, h, hn⟩ := zero + exact ordinaryHypergeometricSeries_eq_zero_of_neg_nat a b c hn h + +theorem ordinaryHypergeometricSeries_norm_div_succ_norm (n : ℕ) + (habc : ∀ kn < n, (↑kn ≠ -a ∧ ↑kn ≠ -b ∧ ↑kn ≠ -c)) : + ‖ordinaryHypergeometricCoefficient a b c n‖ / ‖ordinaryHypergeometricCoefficient a b c n.succ‖ + = ‖a + n‖⁻¹ * ‖b + n‖⁻¹ * ‖c + n‖ * ‖1 + (n : 𝕂)‖ := by + simp only [mul_inv_rev, factorial_succ, cast_mul, cast_add, + cast_one, ascPochhammer_succ_eval, norm_mul, norm_inv] + calc + _ = ‖Polynomial.eval a (ascPochhammer 𝕂 n)‖ * ‖Polynomial.eval a (ascPochhammer 𝕂 n)‖⁻¹ * + ‖Polynomial.eval b (ascPochhammer 𝕂 n)‖ * ‖Polynomial.eval b (ascPochhammer 𝕂 n)‖⁻¹ * + ‖Polynomial.eval c (ascPochhammer 𝕂 n)‖⁻¹⁻¹ * ‖Polynomial.eval c (ascPochhammer 𝕂 n)‖⁻¹ * + ‖(n ! : 𝕂)‖⁻¹⁻¹ * ‖(n ! : 𝕂)‖⁻¹ * ‖a + n‖⁻¹ * ‖b + n‖⁻¹ * ‖c + n‖⁻¹⁻¹ * + ‖1 + (n : 𝕂)‖⁻¹⁻¹ := by ring_nf + _ = _ := by + simp only [inv_inv] + repeat rw [DivisionRing.mul_inv_cancel, one_mul] + all_goals + rw [norm_ne_zero_iff] + any_goals + apply (ascPochhammer_eval_eq_zero_iff n _).not.2 + push_neg + exact fun kn hkn ↦ by simp [habc kn hkn] + exact cast_ne_zero.2 (factorial_ne_zero n) + +/-- The radius of convergence of `ordinaryHypergeometricSeries` is unity if none of the parameters +are non-positive integers. -/ +theorem ordinaryHypergeometricSeries_radius_eq_one + (habc : ∀ kn : ℕ, ↑kn ≠ -a ∧ ↑kn ≠ -b ∧ ↑kn ≠ -c) : + (ordinaryHypergeometricSeries 𝔸 a b c).radius = 1 := by + convert ofScalars_radius_eq_of_tendsto 𝔸 _ one_ne_zero ?_ + suffices Tendsto (fun k : ℕ ↦ (a + k)⁻¹ * (b + k)⁻¹ * (c + k) * ((1 : 𝕂) + k)) atTop (𝓝 1) by + simp_rw [ordinaryHypergeometricSeries_norm_div_succ_norm a b c _ (fun n _ ↦ habc n)] + simp [← norm_mul, ← norm_inv] + convert Filter.Tendsto.norm this + exact norm_one.symm + have (k : ℕ) : (a + k)⁻¹ * (b + k)⁻¹ * (c + k) * ((1 : 𝕂) + k) = + (c + k) / (a + k) * ((1 + k) / (b + k)) := by field_simp + simp_rw [this] + apply (mul_one (1 : 𝕂)) ▸ Filter.Tendsto.mul <;> + convert RCLike.tendsto_add_mul_div_add_mul_atTop_nhds _ _ (1 : 𝕂) one_ne_zero <;> simp + +end RCLike diff --git a/Mathlib/Analysis/SpecificLimits/RCLike.lean b/Mathlib/Analysis/SpecificLimits/RCLike.lean index 287d854e7d39d..2f82b7954f882 100644 --- a/Mathlib/Analysis/SpecificLimits/RCLike.lean +++ b/Mathlib/Analysis/SpecificLimits/RCLike.lean @@ -22,3 +22,18 @@ theorem RCLike.tendsto_inverse_atTop_nhds_zero_nat : simp @[deprecated (since := "2024-01-16")] alias RCLike.tendsto_inverse_atTop_nhds_0_nat := RCLike.tendsto_inverse_atTop_nhds_zero_nat + +variable {𝕜} + +theorem RCLike.tendsto_add_mul_div_add_mul_atTop_nhds (a b c : 𝕜) {d : 𝕜} (hd : d ≠ 0) : + Tendsto (fun k : ℕ ↦ (a + c * k) / (b + d * k)) atTop (𝓝 (c / d)) := by + apply Filter.Tendsto.congr' + case f₁ => exact fun k ↦ (a * (↑k)⁻¹ + c) / (b * (↑k)⁻¹ + d) + refine (eventually_ne_atTop 0).mp (Eventually.of_forall ?_) + · intro h hx + field_simp [hx] + · apply Filter.Tendsto.div _ _ hd + all_goals + apply zero_add (_ : 𝕜) ▸ Filter.Tendsto.add_const _ _ + apply mul_zero (_ : 𝕜) ▸ Filter.Tendsto.const_mul _ _ + exact RCLike.tendsto_inverse_atTop_nhds_zero_nat 𝕜 diff --git a/Mathlib/RingTheory/Polynomial/Pochhammer.lean b/Mathlib/RingTheory/Polynomial/Pochhammer.lean index a96e49ebd3495..fcb2365cac7b2 100644 --- a/Mathlib/RingTheory/Polynomial/Pochhammer.lean +++ b/Mathlib/RingTheory/Polynomial/Pochhammer.lean @@ -3,9 +3,7 @@ Copyright (c) 2020 Kim Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison -/ -import Mathlib.Algebra.Polynomial.Degree.Definitions -import Mathlib.Algebra.Polynomial.Eval -import Mathlib.Algebra.Polynomial.Monic +import Mathlib.Algebra.CharP.Defs import Mathlib.Algebra.Polynomial.RingDivision import Mathlib.Tactic.Abel @@ -24,6 +22,8 @@ that are focused on `Nat` can be found in `Data.Nat.Factorial` as `Nat.ascFactor As with many other families of polynomials, even though the coefficients are always in `ℕ` or `ℤ` , we define the polynomial with coefficients in any `[Semiring S]` or `[Ring R]`. +In an integral domain `S`, we show that `ascPochhammer S n` is zero iff +`n` is a sufficiently large non-positive integer. ## TODO @@ -372,4 +372,38 @@ theorem descPochhammer_int_eq_ascFactorial (a b : ℕ) : rw [← Nat.cast_add, descPochhammer_eval_eq_descFactorial ℤ (a + b) b, Nat.add_descFactorial_eq_ascFactorial] +variable {R} + +/-- The Pochhammer polynomial of degree `n` has roots at `0`, `-1`, ..., `-(n - 1)`. -/ +theorem ascPochhammer_eval_neg_coe_nat_of_lt {n k : ℕ} (h : k < n) : + (ascPochhammer R n).eval (-(k : R)) = 0 := by + induction n with + | zero => contradiction + | succ n ih => + rw [ascPochhammer_succ_eval] + rcases lt_trichotomy k n with hkn | rfl | hkn + · simp [ih hkn] + · simp + · omega + +/-- Over an integral domain, the Pochhammer polynomial of degree `n` has roots *only* at +`0`, `-1`, ..., `-(n - 1)`. -/ +@[simp] +theorem ascPochhammer_eval_eq_zero_iff [IsDomain R] + (n : ℕ) (r : R) : (ascPochhammer R n).eval r = 0 ↔ ∃ k < n, k = -r := by + refine ⟨fun zero' ↦ ?_, fun hrn ↦ ?_⟩ + · induction n with + | zero => simp only [ascPochhammer_zero, Polynomial.eval_one, one_ne_zero] at zero' + | succ n ih => + rw [ascPochhammer_succ_eval, mul_eq_zero] at zero' + cases zero' with + | inl h => + obtain ⟨rn, hrn, rrn⟩ := ih h + exact ⟨rn, by omega, rrn⟩ + | inr h => + exact ⟨n, lt_add_one n, eq_neg_of_add_eq_zero_right h⟩ + · obtain ⟨rn, hrn, rnn⟩ := hrn + convert ascPochhammer_eval_neg_coe_nat_of_lt hrn + simp [rnn] + end Ring From c41f54d5da1cfb4821b762dadca0723eaf0de39e Mon Sep 17 00:00:00 2001 From: Oliver Nash Date: Thu, 17 Oct 2024 08:31:04 +0000 Subject: [PATCH 059/505] feat: generalise `LieModule.independent_genWeightSpace` to any family of compatible endomorphisms (#17654) Also create new file to group together results about simultaneous (generalised) eigenspaces. Co-authored-by: Oliver Nash <7734364+ocfnash@users.noreply.github.com> --- Mathlib.lean | 1 + Mathlib/Algebra/Lie/Weights/Basic.lean | 49 +---- Mathlib/LinearAlgebra/Eigenspace/Basic.lean | 32 --- Mathlib/LinearAlgebra/Eigenspace/Pi.lean | 193 ++++++++++++++++++ .../Eigenspace/Triangularizable.lean | 51 ----- 5 files changed, 200 insertions(+), 126 deletions(-) create mode 100644 Mathlib/LinearAlgebra/Eigenspace/Pi.lean diff --git a/Mathlib.lean b/Mathlib.lean index 1170098fd659e..1595a1b1c093b 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3033,6 +3033,7 @@ import Mathlib.LinearAlgebra.Dual import Mathlib.LinearAlgebra.Eigenspace.Basic import Mathlib.LinearAlgebra.Eigenspace.Matrix import Mathlib.LinearAlgebra.Eigenspace.Minpoly +import Mathlib.LinearAlgebra.Eigenspace.Pi import Mathlib.LinearAlgebra.Eigenspace.Semisimple import Mathlib.LinearAlgebra.Eigenspace.Triangularizable import Mathlib.LinearAlgebra.Eigenspace.Zero diff --git a/Mathlib/Algebra/Lie/Weights/Basic.lean b/Mathlib/Algebra/Lie/Weights/Basic.lean index e0f030da4cacc..1d5d4302ac426 100644 --- a/Mathlib/Algebra/Lie/Weights/Basic.lean +++ b/Mathlib/Algebra/Lie/Weights/Basic.lean @@ -6,7 +6,7 @@ Authors: Oliver Nash import Mathlib.Algebra.Ring.Divisibility.Lemmas import Mathlib.Algebra.Lie.Nilpotent import Mathlib.Algebra.Lie.Engel -import Mathlib.LinearAlgebra.Eigenspace.Triangularizable +import Mathlib.LinearAlgebra.Eigenspace.Pi import Mathlib.RingTheory.Artinian import Mathlib.LinearAlgebra.Trace import Mathlib.LinearAlgebra.FreeModule.PID @@ -650,48 +650,11 @@ lemma injOn_genWeightSpace [NoZeroSMulDivisors R M] : See also `LieModule.independent_genWeightSpace'`. -/ lemma independent_genWeightSpace [NoZeroSMulDivisors R M] : - CompleteLattice.Independent fun (χ : L → R) ↦ genWeightSpace M χ := by - classical - suffices ∀ χ (s : Finset (L → R)) (_ : χ ∉ s), - Disjoint (genWeightSpace M χ) (s.sup fun (χ : L → R) ↦ genWeightSpace M χ) by - simpa only [CompleteLattice.independent_iff_supIndep_of_injOn (injOn_genWeightSpace R L M), - Finset.supIndep_iff_disjoint_erase] using fun s χ _ ↦ this _ _ (s.not_mem_erase χ) - intro χ₁ s - induction s using Finset.induction_on with - | empty => simp - | insert _n ih => - rename_i χ₂ s - intro hχ₁₂ - obtain ⟨hχ₁₂ : χ₁ ≠ χ₂, hχ₁ : χ₁ ∉ s⟩ := by rwa [Finset.mem_insert, not_or] at hχ₁₂ - specialize ih hχ₁ - rw [Finset.sup_insert, disjoint_iff, LieSubmodule.eq_bot_iff] - rintro x ⟨hx, hx'⟩ - simp only [SetLike.mem_coe, LieSubmodule.mem_coeSubmodule] at hx hx' - suffices x ∈ genWeightSpace M χ₂ by - rw [← LieSubmodule.mem_bot (R := R) (L := L), ← (disjoint_genWeightSpace R L M hχ₁₂).eq_bot] - exact ⟨hx, this⟩ - obtain ⟨y, hy, z, hz, rfl⟩ := (LieSubmodule.mem_sup _ _ _).mp hx'; clear hx' - suffices ∀ l, ∃ (k : ℕ), - ((toEnd R L M l - algebraMap R (Module.End R M) (χ₂ l)) ^ k) (y + z) ∈ - genWeightSpace M χ₁ ⊓ Finset.sup s fun χ ↦ genWeightSpace M χ by - simpa only [ih.eq_bot, LieSubmodule.mem_bot, mem_genWeightSpace] using this - intro l - let g : Module.End R M := toEnd R L M l - algebraMap R (Module.End R M) (χ₂ l) - obtain ⟨k, hk : (g ^ k) y = 0⟩ := (mem_genWeightSpace _ _ _).mp hy l - refine ⟨k, (LieSubmodule.mem_inf _ _ _).mp ⟨?_, ?_⟩⟩ - · exact LieSubmodule.mapsTo_pow_toEnd_sub_algebraMap _ hx - · rw [map_add, hk, zero_add] - suffices (s.sup fun χ ↦ genWeightSpace M χ : Submodule R M).map (g ^ k) ≤ - s.sup fun χ ↦ genWeightSpace M χ by - refine this (Submodule.mem_map_of_mem ?_) - simp_rw [← LieSubmodule.mem_coeSubmodule, Finset.sup_eq_iSup, - LieSubmodule.iSup_coe_toSubmodule, ← Finset.sup_eq_iSup] at hz - exact hz - simp_rw [Finset.sup_eq_iSup, Submodule.map_iSup (ι := L → R), Submodule.map_iSup (ι := _ ∈ s), - LieSubmodule.iSup_coe_toSubmodule] - refine iSup₂_mono fun χ _ ↦ ?_ - rintro - ⟨u, hu, rfl⟩ - exact LieSubmodule.mapsTo_pow_toEnd_sub_algebraMap _ hu + CompleteLattice.Independent fun χ : L → R ↦ genWeightSpace M χ := by + simp only [LieSubmodule.independent_iff_coe_toSubmodule, genWeightSpace, + LieSubmodule.iInf_coe_toSubmodule] + exact Module.End.independent_iInf_maxGenEigenspace_of_forall_mapsTo (toEnd R L M) + (fun x y φ z ↦ (genWeightSpaceOf M φ y).lie_mem) lemma independent_genWeightSpace' [NoZeroSMulDivisors R M] : CompleteLattice.Independent fun χ : Weight R L M ↦ genWeightSpace M χ := diff --git a/Mathlib/LinearAlgebra/Eigenspace/Basic.lean b/Mathlib/LinearAlgebra/Eigenspace/Basic.lean index e2e7aa9fcb4c9..a66604891270f 100644 --- a/Mathlib/LinearAlgebra/Eigenspace/Basic.lean +++ b/Mathlib/LinearAlgebra/Eigenspace/Basic.lean @@ -787,38 +787,6 @@ lemma _root_.Submodule.inf_genEigenspace (f : End R M) (p : Submodule R M) {k : (genEigenspace (LinearMap.restrict f hfp) μ k).map p.subtype := by rw [f.genEigenspace_restrict _ _ _ hfp, Submodule.map_comap_eq, Submodule.range_subtype] -/-- Given a family of endomorphisms `i ↦ f i`, a family of candidate eigenvalues `i ↦ μ i`, and a -submodule `p` which is invariant wrt every `f i`, the intersection of `p` with the simultaneous -maximal generalised eigenspace (taken over all `i`), is the same as the simultaneous maximal -generalised eigenspace of the `f i` restricted to `p`. -/ -lemma _root_.Submodule.inf_iInf_maxGenEigenspace_of_forall_mapsTo {ι : Type*} {μ : ι → R} - (f : ι → End R M) (p : Submodule R M) (hfp : ∀ i, MapsTo (f i) p p) : - p ⊓ ⨅ i, (f i).maxGenEigenspace (μ i) = - (⨅ 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_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 -distinguished index `i` whose maximal generalised `μ i`-eigenspace is invariant wrt every `f j`, -taking simultaneous maximal generalised eigenspaces is unaffected by first restricting to the -distinguished generalised `μ i`-eigenspace. -/ -lemma iInf_maxGenEigenspace_restrict_map_subtype_eq - {ι : Type*} {μ : ι → R} (i : ι) (f : ι → End R M) - (h : ∀ j, MapsTo (f j) ((f i).maxGenEigenspace (μ i)) ((f i).maxGenEigenspace (μ i))) : - letI p := (f i).maxGenEigenspace (μ i) - letI q (j : ι) := maxGenEigenspace ((f j).restrict (h j)) (μ j) - (⨅ j, q j).map p.subtype = ⨅ j, (f j).maxGenEigenspace (μ j) := by - have : Nonempty ι := ⟨i⟩ - set p := (f i).maxGenEigenspace (μ i) - have : ⨅ j, (f j).maxGenEigenspace (μ j) = p ⊓ ⨅ j, (f j).maxGenEigenspace (μ j) := by - refine le_antisymm ?_ inf_le_right - simpa only [le_inf_iff, le_refl, and_true] using iInf_le _ _ - 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 _)] - 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} (h : MapsTo f (g.maxGenEigenspace μ₁) (g.maxGenEigenspace μ₂)) : diff --git a/Mathlib/LinearAlgebra/Eigenspace/Pi.lean b/Mathlib/LinearAlgebra/Eigenspace/Pi.lean new file mode 100644 index 0000000000000..3eeacc0153eb7 --- /dev/null +++ b/Mathlib/LinearAlgebra/Eigenspace/Pi.lean @@ -0,0 +1,193 @@ +/- +Copyright (c) 2024 Oliver Nash. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Oliver Nash +-/ +import Mathlib.LinearAlgebra.Eigenspace.Triangularizable + +/-! +# Simultaneous eigenvectors and eigenvalues for families of endomorphisms + +In finite dimensions, the theory of simultaneous eigenvalues for a family of linear endomorphisms +`i ↦ f i` enjoys similar properties to that of a single endomorphism, provided the family obeys a +compatibilty condition. This condition is that the maximum generalised eigenspaces of each +endomorphism are invariant under the action of all members of the family. It is trivially satisfied +for commuting endomorphisms but there are important more general situations where it also holds +(e.g., representations of nilpotent Lie algebras). + +## Main definitions / results + * `Module.End.independent_iInf_maxGenEigenspace_of_forall_mapsTo`: the simultaneous generalised + eigenspaces of a compatible family of endomorphisms are independent. + * `Module.End.iSup_iInf_maxGenEigenspace_eq_top_of_forall_mapsTo`: in finite dimensions, the + simultaneous generalised eigenspaces of a compatible family of endomorphisms span if the same + is true of each map individually. + +-/ + +open Function Set + +namespace Module.End + +variable {ι R K M : Type*} [CommRing R] [Field K] [AddCommGroup M] [Module R M] [Module K M] + (f : ι → End R M) + +theorem mem_iInf_maxGenEigenspace_iff (χ : ι → R) (m : M) : + m ∈ ⨅ i, (f i).maxGenEigenspace (χ i) ↔ ∀ j, ∃ k : ℕ, ((f j - χ j • ↑1) ^ k) m = 0 := by + simp + +/-- Given a family of endomorphisms `i ↦ f i`, a family of candidate eigenvalues `i ↦ μ i`, and a +submodule `p` which is invariant wrt every `f i`, the intersection of `p` with the simultaneous +maximal generalised eigenspace (taken over all `i`), is the same as the simultaneous maximal +generalised eigenspace of the `f i` restricted to `p`. -/ +lemma _root_.Submodule.inf_iInf_maxGenEigenspace_of_forall_mapsTo {μ : ι → R} + (p : Submodule R M) (hfp : ∀ i, MapsTo (f i) p p) : + p ⊓ ⨅ i, (f i).maxGenEigenspace (μ i) = + (⨅ 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_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 +distinguished index `i` whose maximal generalised `μ i`-eigenspace is invariant wrt every `f j`, +taking simultaneous maximal generalised eigenspaces is unaffected by first restricting to the +distinguished generalised `μ i`-eigenspace. -/ +lemma iInf_maxGenEigenspace_restrict_map_subtype_eq + {μ : ι → R} (i : ι) + (h : ∀ j, MapsTo (f j) ((f i).maxGenEigenspace (μ i)) ((f i).maxGenEigenspace (μ i))) : + letI p := (f i).maxGenEigenspace (μ i) + letI q (j : ι) := maxGenEigenspace ((f j).restrict (h j)) (μ j) + (⨅ j, q j).map p.subtype = ⨅ j, (f j).maxGenEigenspace (μ j) := by + have : Nonempty ι := ⟨i⟩ + set p := (f i).maxGenEigenspace (μ i) + have : ⨅ j, (f j).maxGenEigenspace (μ j) = p ⊓ ⨅ j, (f j).maxGenEigenspace (μ j) := by + refine le_antisymm ?_ inf_le_right + simpa only [le_inf_iff, le_refl, and_true] using iInf_le _ _ + 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 _)] + +variable [NoZeroSMulDivisors R M] + +lemma disjoint_iInf_maxGenEigenspace {χ₁ χ₂ : ι → R} (h : χ₁ ≠ χ₂) : + Disjoint (⨅ i, (f i).maxGenEigenspace (χ₁ i)) (⨅ i, (f i).maxGenEigenspace (χ₂ i)) := by + obtain ⟨j, hj⟩ : ∃ j, χ₁ j ≠ χ₂ j := Function.ne_iff.mp h + exact (End.disjoint_unifEigenspace (f j) hj ⊤ ⊤).mono (iInf_le _ j) (iInf_le _ j) + +lemma injOn_iInf_maxGenEigenspace : + InjOn (fun χ : ι → R ↦ ⨅ i, (f i).maxGenEigenspace (χ i)) + {χ | ⨅ i, (f i).maxGenEigenspace (χ i) ≠ ⊥} := by + rintro χ₁ _ χ₂ + hχ₂ (hχ₁₂ : ⨅ i, (f i).maxGenEigenspace (χ₁ i) = ⨅ i, (f i).maxGenEigenspace (χ₂ i)) + contrapose! hχ₂ + simpa [hχ₁₂] using disjoint_iInf_maxGenEigenspace f hχ₂ + +lemma independent_iInf_maxGenEigenspace_of_forall_mapsTo + (h : ∀ i j φ, MapsTo (f i) ((f j).maxGenEigenspace φ) ((f j).maxGenEigenspace φ)) : + CompleteLattice.Independent fun χ : ι → R ↦ ⨅ i, (f i).maxGenEigenspace (χ i) := by + replace h (l : ι) (χ : ι → R) : + MapsTo (f l) (⨅ i, (f i).maxGenEigenspace (χ i)) (⨅ i, (f i).maxGenEigenspace (χ i)) := by + intro x hx + simp only [iInf_eq_iInter, mem_iInter, SetLike.mem_coe] at hx ⊢ + exact fun i ↦ h l i (χ i) (hx i) + classical + suffices ∀ χ (s : Finset (ι → R)) (_ : χ ∉ s), + Disjoint (⨅ i, (f i).maxGenEigenspace (χ i)) + (s.sup fun (χ : ι → R) ↦ ⨅ i, (f i).maxGenEigenspace (χ i)) by + simpa only [CompleteLattice.independent_iff_supIndep_of_injOn (injOn_iInf_maxGenEigenspace f), + Finset.supIndep_iff_disjoint_erase] using fun s χ _ ↦ this _ _ (s.not_mem_erase χ) + intro χ₁ s + induction s using Finset.induction_on with + | empty => simp + | insert _n ih => + rename_i χ₂ s + intro hχ₁₂ + obtain ⟨hχ₁₂ : χ₁ ≠ χ₂, hχ₁ : χ₁ ∉ s⟩ := by rwa [Finset.mem_insert, not_or] at hχ₁₂ + specialize ih hχ₁ + rw [Finset.sup_insert, disjoint_iff, Submodule.eq_bot_iff] + rintro x ⟨hx, hx'⟩ + simp only [SetLike.mem_coe] at hx hx' + suffices x ∈ ⨅ i, (f i).maxGenEigenspace (χ₂ i) by + rw [← Submodule.mem_bot (R := R), ← (disjoint_iInf_maxGenEigenspace f hχ₁₂).eq_bot] + exact ⟨hx, this⟩ + obtain ⟨y, hy, z, hz, rfl⟩ := Submodule.mem_sup.mp hx'; clear hx' + suffices ∀ l, ∃ (k : ℕ), + ((f l - algebraMap R (Module.End R M) (χ₂ l)) ^ k) (y + z) ∈ + (⨅ i, (f i).maxGenEigenspace (χ₁ i)) ⊓ + Finset.sup s fun χ ↦ ⨅ i, (f i).maxGenEigenspace (χ i) by + simpa [ih.eq_bot, Submodule.mem_bot] using this + intro l + let g : Module.End R M := f l - algebraMap R (Module.End R M) (χ₂ l) + obtain ⟨k, hk : (g ^ k) y = 0⟩ := (mem_iInf_maxGenEigenspace_iff _ _ _).mp hy l + have aux (f : End R M) (φ : R) (k : ℕ) (p : Submodule R M) (hp : MapsTo f p p) : + MapsTo ((f - algebraMap R (Module.End R M) φ) ^ k) p p := by + rw [LinearMap.coe_pow] + exact MapsTo.iterate (fun m hm ↦ p.sub_mem (hp hm) (p.smul_mem _ hm)) k + refine ⟨k, Submodule.mem_inf.mp ⟨?_, ?_⟩⟩ + · refine aux (f l) (χ₂ l) k (⨅ i, (f i).maxGenEigenspace (χ₁ i)) ?_ hx + simp only [Submodule.iInf_coe] + exact h l χ₁ + · rw [map_add, hk, zero_add] + suffices (s.sup fun χ ↦ (⨅ i, (f i).maxGenEigenspace (χ i))).map (g ^ k) ≤ + s.sup fun χ ↦ (⨅ i, (f i).maxGenEigenspace (χ i)) by + refine this (Submodule.mem_map_of_mem ?_) + simp_rw [Finset.sup_eq_iSup, ← Finset.sup_eq_iSup] at hz + exact hz + simp_rw [Finset.sup_eq_iSup, Submodule.map_iSup (ι := ι → R), Submodule.map_iSup (ι := _ ∈ s)] + refine iSup₂_mono fun χ _ ↦ ?_ + rintro - ⟨u, hu, rfl⟩ + refine aux (f l) (χ₂ l) k (⨅ i, (f i).maxGenEigenspace (χ i)) ?_ hu + simp only [Submodule.iInf_coe] + exact h l χ + +/-- Given a family of endomorphisms `i ↦ f i` which are compatible in the sense that every maximal +generalised eigenspace of `f i` is invariant wrt `f j`, if each `f i` is triangularizable, the +family is simultaneously triangularizable. -/ +lemma iSup_iInf_maxGenEigenspace_eq_top_of_forall_mapsTo [FiniteDimensional K M] + (f : ι → End K M) + (h : ∀ i j φ, MapsTo (f i) ((f j).maxGenEigenspace φ) ((f j).maxGenEigenspace φ)) + (h' : ∀ i, ⨆ μ, (f i).maxGenEigenspace μ = ⊤) : + ⨆ χ : ι → K, ⨅ i, (f i).maxGenEigenspace (χ i) = ⊤ := by + generalize h_dim : finrank K M = n + induction n using Nat.strongRecOn generalizing M with | ind n ih => ?_ + obtain this | ⟨i : ι, hy : ¬ ∃ φ, (f i).maxGenEigenspace φ = ⊤⟩ := + forall_or_exists_not (fun j : ι ↦ ∃ φ : K, (f j).maxGenEigenspace φ = ⊤) + · choose χ hχ using this + replace hχ : ⨅ i, (f i).maxGenEigenspace (χ i) = ⊤ := by simpa + simp_rw [eq_top_iff] at hχ ⊢ + exact le_trans hχ <| le_iSup (fun χ : ι → K ↦ ⨅ i, (f i).maxGenEigenspace (χ i)) χ + · replace hy : ∀ φ, finrank K ((f i).maxGenEigenspace φ) < n := fun φ ↦ by + simp_rw [not_exists, ← lt_top_iff_ne_top] at hy; exact h_dim ▸ Submodule.finrank_lt (hy φ) + have hi (j : ι) (φ : K) : + MapsTo (f j) ((f i).maxGenEigenspace φ) ((f i).maxGenEigenspace φ) := by + exact h j i φ + replace ih (φ : K) : + ⨆ χ : ι → K, ⨅ j, maxGenEigenspace ((f j).restrict (hi j φ)) (χ j) = ⊤ := by + apply ih _ (hy φ) + · intro j k μ + exact mapsTo_restrict_maxGenEigenspace_restrict_of_mapsTo (f j) (f k) _ _ (h j k μ) + · 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 + suffices ∀ χ : ι → K, χ i ≠ φ → ⨅ j, maxGenEigenspace ((f j).restrict (hi j φ)) (χ j) = ⊥ by + specialize ih φ; rw [iSup_split, biSup_congr this] at ih; simpa using ih + intro χ hχ + 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) = + maxGenEigenspace (f i) φ := by + have (χ : ι → K) (hχ : χ i = φ) : ⨅ j, maxGenEigenspace (f j) (χ j) = + (⨅ j, maxGenEigenspace ((f j).restrict (hi j φ)) (χ j)).map + ((f i).maxGenEigenspace φ).subtype := by + rw [← hχ, iInf_maxGenEigenspace_restrict_map_subtype_eq] + simp_rw [biSup_congr this, ← Submodule.map_iSup, ih, Submodule.map_top, + Submodule.range_subtype] + simpa only [← ih, iSup_comm (ι := K), iSup_iSup_eq_right] using h' i + +end Module.End diff --git a/Mathlib/LinearAlgebra/Eigenspace/Triangularizable.lean b/Mathlib/LinearAlgebra/Eigenspace/Triangularizable.lean index 03e4ac0132b66..e619e16819564 100644 --- a/Mathlib/LinearAlgebra/Eigenspace/Triangularizable.lean +++ b/Mathlib/LinearAlgebra/Eigenspace/Triangularizable.lean @@ -222,54 +222,3 @@ theorem Module.End.iSup_genEigenspace_restrict_eq_top simp_rw [Submodule.inf_genEigenspace f p h, Submodule.comap_subtype_self, ← Submodule.map_iSup, Submodule.comap_map_eq_of_injective h_inj] at this exact this.symm - -/-- Given a family of endomorphisms `i ↦ f i` which are compatible in the sense that every maximal -generalised eigenspace of `f i` is invariant wrt `f j`, if each `f i` is triangularizable, the -family is simultaneously triangularizable. -/ -lemma Module.End.iSup_iInf_maxGenEigenspace_eq_top_of_forall_mapsTo - {ι : Type*} [FiniteDimensional K V] - (f : ι → End K V) - (h : ∀ i j φ, MapsTo (f i) ((f j).maxGenEigenspace φ) ((f j).maxGenEigenspace φ)) - (h' : ∀ i, ⨆ μ, (f i).maxGenEigenspace μ = ⊤) : - ⨆ χ : ι → K, ⨅ i, (f i).maxGenEigenspace (χ i) = ⊤ := by - generalize h_dim : finrank K V = n - induction n using Nat.strongRecOn generalizing V with | ind n ih => ?_ - obtain this | ⟨i : ι, hy : ¬ ∃ φ, (f i).maxGenEigenspace φ = ⊤⟩ := - forall_or_exists_not (fun j : ι ↦ ∃ φ : K, (f j).maxGenEigenspace φ = ⊤) - · choose χ hχ using this - replace hχ : ⨅ i, (f i).maxGenEigenspace (χ i) = ⊤ := by simpa - simp_rw [eq_top_iff] at hχ ⊢ - exact le_trans hχ <| le_iSup (fun χ : ι → K ↦ ⨅ i, (f i).maxGenEigenspace (χ i)) χ - · replace hy : ∀ φ, finrank K ((f i).maxGenEigenspace φ) < n := fun φ ↦ by - simp_rw [not_exists, ← lt_top_iff_ne_top] at hy; exact h_dim ▸ Submodule.finrank_lt (hy φ) - have hi (j : ι) (φ : K) : - MapsTo (f j) ((f i).maxGenEigenspace φ) ((f i).maxGenEigenspace φ) := by - exact h j i φ - replace ih (φ : K) : - ⨆ χ : ι → K, ⨅ j, maxGenEigenspace ((f j).restrict (hi j φ)) (χ j) = ⊤ := by - apply ih _ (hy φ) - · intro j k μ - exact mapsTo_restrict_maxGenEigenspace_restrict_of_mapsTo (f j) (f k) _ _ (h j k μ) - · 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 - suffices ∀ χ : ι → K, χ i ≠ φ → ⨅ j, maxGenEigenspace ((f j).restrict (hi j φ)) (χ j) = ⊥ by - specialize ih φ; rw [iSup_split, biSup_congr this] at ih; simpa using ih - intro χ hχ - 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) = - maxGenEigenspace (f i) φ := by - have (χ : ι → K) (hχ : χ i = φ) : ⨅ j, maxGenEigenspace (f j) (χ j) = - (⨅ j, maxGenEigenspace ((f j).restrict (hi j φ)) (χ j)).map - ((f i).maxGenEigenspace φ).subtype := by - rw [← hχ, iInf_maxGenEigenspace_restrict_map_subtype_eq] - simp_rw [biSup_congr this, ← Submodule.map_iSup, ih, Submodule.map_top, - Submodule.range_subtype] - simpa only [← ih, iSup_comm (ι := K), iSup_iSup_eq_right] using h' i From b1ec18a85983bb1edd81b603b996cd9f9f7ee8df Mon Sep 17 00:00:00 2001 From: Christian Merten Date: Thu, 17 Oct 2024 08:31:05 +0000 Subject: [PATCH 060/505] feat(CategoryTheory/Sites): pretopology induced by a morphism property (#17736) Introduces the pretopology induced by a morphism property satisfying sufficient stability properties. --- Mathlib.lean | 1 + .../MorphismProperty/Composition.lean | 20 ++++++ .../MorphismProperty/Limits.lean | 10 +++ .../Sites/MorphismProperty.lean | 66 +++++++++++++++++++ 4 files changed, 97 insertions(+) create mode 100644 Mathlib/CategoryTheory/Sites/MorphismProperty.lean diff --git a/Mathlib.lean b/Mathlib.lean index 1595a1b1c093b..d978b6190e81c 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1905,6 +1905,7 @@ import Mathlib.CategoryTheory.Sites.LocallyFullyFaithful import Mathlib.CategoryTheory.Sites.LocallyInjective import Mathlib.CategoryTheory.Sites.LocallySurjective import Mathlib.CategoryTheory.Sites.MayerVietorisSquare +import Mathlib.CategoryTheory.Sites.MorphismProperty import Mathlib.CategoryTheory.Sites.NonabelianCohomology.H1 import Mathlib.CategoryTheory.Sites.OneHypercover import Mathlib.CategoryTheory.Sites.Over diff --git a/Mathlib/CategoryTheory/MorphismProperty/Composition.lean b/Mathlib/CategoryTheory/MorphismProperty/Composition.lean index aff2f87e7acc0..13fd8d28c5d30 100644 --- a/Mathlib/CategoryTheory/MorphismProperty/Composition.lean +++ b/Mathlib/CategoryTheory/MorphismProperty/Composition.lean @@ -49,6 +49,10 @@ instance inverseImage {P : MorphismProperty D} [P.ContainsIdentities] (F : C ⥤ (P.inverseImage F).ContainsIdentities where id_mem X := by simpa only [← F.map_id] using P.id_mem (F.obj X) +instance inf {P Q : MorphismProperty C} [P.ContainsIdentities] [Q.ContainsIdentities] : + (P ⊓ Q).ContainsIdentities where + id_mem X := ⟨P.id_mem X, Q.id_mem X⟩ + end ContainsIdentities instance Prod.containsIdentities {C₁ C₂ : Type*} [Category C₁] [Category C₂] @@ -61,6 +65,14 @@ instance Pi.containsIdentities {J : Type w} {C : J → Type u} (pi W).ContainsIdentities := ⟨fun _ _ => MorphismProperty.id_mem _ _⟩ +lemma of_isIso (P : MorphismProperty C) [P.ContainsIdentities] [P.RespectsIso] {X Y : C} (f : X ⟶ Y) + [IsIso f] : P f := + Category.id_comp f ▸ RespectsIso.postcomp P f (𝟙 X) (P.id_mem X) + +lemma isomorphisms_le_of_containsIdentities (P : MorphismProperty C) [P.ContainsIdentities] + [P.RespectsIso] : + isomorphisms C ≤ P := fun _ _ f (_ : IsIso f) ↦ P.of_isIso f + /-- A morphism property satisfies `IsStableUnderComposition` if the composition of two such morphisms still falls in the class. -/ class IsStableUnderComposition (P : MorphismProperty C) : Prop where @@ -78,6 +90,11 @@ instance IsStableUnderComposition.unop {P : MorphismProperty Cᵒᵖ} [P.IsStabl P.unop.IsStableUnderComposition where comp_mem f g hf hg := P.comp_mem g.op f.op hg hf +instance IsStableUnderComposition.inf {P Q : MorphismProperty C} [P.IsStableUnderComposition] + [Q.IsStableUnderComposition] : + (P ⊓ Q).IsStableUnderComposition where + comp_mem f g hf hg := ⟨P.comp_mem f g hf.left hg.left, Q.comp_mem f g hf.right hg.right⟩ + /-- A morphism property is `StableUnderInverse` if the inverse of a morphism satisfying the property still falls in the class. -/ def StableUnderInverse (P : MorphismProperty C) : Prop := @@ -167,6 +184,9 @@ instance : (epimorphisms C).IsMultiplicative where instance {P : MorphismProperty D} [P.IsMultiplicative] (F : C ⥤ D) : (P.inverseImage F).IsMultiplicative where +instance inf {P Q : MorphismProperty C} [P.IsMultiplicative] [Q.IsMultiplicative] : + (P ⊓ Q).IsMultiplicative where + end IsMultiplicative /-- A class of morphisms `W` has the two-out-of-three property if whenever two out diff --git a/Mathlib/CategoryTheory/MorphismProperty/Limits.lean b/Mathlib/CategoryTheory/MorphismProperty/Limits.lean index 7352e3fdd289a..97d95f0510ebd 100644 --- a/Mathlib/CategoryTheory/MorphismProperty/Limits.lean +++ b/Mathlib/CategoryTheory/MorphismProperty/Limits.lean @@ -163,6 +163,16 @@ theorem StableUnderBaseChange.op {P : MorphismProperty C} (hP : StableUnderBaseC theorem StableUnderBaseChange.unop {P : MorphismProperty Cᵒᵖ} (hP : StableUnderBaseChange P) : StableUnderCobaseChange P.unop := fun _ _ _ _ _ _ _ _ sq hf => hP sq.op hf +lemma StableUnderBaseChange.inf {P Q : MorphismProperty C} (hP : StableUnderBaseChange P) + (hQ : StableUnderBaseChange Q) : + StableUnderBaseChange (P ⊓ Q) := + fun _ _ _ _ _ _ _ _ hp hg ↦ ⟨hP hp hg.left, hQ hp hg.right⟩ + +lemma StableUnderCobaseChange.inf {P Q : MorphismProperty C} (hP : StableUnderCobaseChange P) + (hQ : StableUnderCobaseChange Q) : + StableUnderCobaseChange (P ⊓ Q) := + fun _ _ _ _ _ _ _ _ hp hg ↦ ⟨hP hp hg.left, hQ hp hg.right⟩ + section variable (W : MorphismProperty C) diff --git a/Mathlib/CategoryTheory/Sites/MorphismProperty.lean b/Mathlib/CategoryTheory/Sites/MorphismProperty.lean new file mode 100644 index 0000000000000..d840eb5deaf0f --- /dev/null +++ b/Mathlib/CategoryTheory/Sites/MorphismProperty.lean @@ -0,0 +1,66 @@ +/- +Copyright (c) 2024 Christian Merten. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Christian Merten +-/ +import Mathlib.CategoryTheory.MorphismProperty.Limits +import Mathlib.CategoryTheory.Sites.Pretopology + +/-! +# The site induced by a morphism property + +Let `C` be a category with pullbacks and `P` be a multiplicative morphism property which is +stable under base change. Then `P` induces a pretopology, where coverings are given by presieves +whose elements satisfy `P`. + +Standard examples of pretopologies in algebraic geometry, such as the étale site, are obtained from +this construction by intersecting with the pretopology of surjective families. + +-/ + +namespace CategoryTheory + +open Limits + +variable {C : Type*} [Category C] [HasPullbacks C] + +namespace MorphismProperty + +/-- If `P` is a multiplicative morphism property which is stable under base change on a category +`C` with pullbacks, then `P` induces a pretopology, where coverings are given by presieves whose +elements satisfy `P`. -/ +def pretopology (P : MorphismProperty C) [P.IsMultiplicative] (hPb : P.StableUnderBaseChange) : + Pretopology C where + coverings X S := ∀ {Y : C} {f : Y ⟶ X}, S f → P f + has_isos X Y f h Z g hg := by + cases hg + haveI : P.RespectsIso := hPb.respectsIso + exact P.of_isIso f + pullbacks X Y f S hS Z g hg := by + obtain ⟨Z, g, hg⟩ := hg + apply hPb.snd g f (hS hg) + transitive X S Ti hS hTi Y f hf := by + obtain ⟨Z, g, h, H, H', rfl⟩ := hf + exact comp_mem _ _ _ (hTi h H H') (hS H) + +/-- To a morphism property `P` satisfying the conditions of `MorphismProperty.pretopology`, we +associate the Grothendieck topology generated by `P.pretopology`. -/ +abbrev grothendieckTopology (P : MorphismProperty C) [P.IsMultiplicative] + (hPb : P.StableUnderBaseChange) : GrothendieckTopology C := + (P.pretopology hPb).toGrothendieck + +variable {P Q : MorphismProperty C} + [P.IsMultiplicative] (hPb : P.StableUnderBaseChange) + [Q.IsMultiplicative] (hQb : Q.StableUnderBaseChange) + +lemma pretopology_le (hPQ : P ≤ Q) : P.pretopology hPb ≤ Q.pretopology hQb := + fun _ _ hS _ f hf ↦ hPQ f (hS hf) + +variable (P Q) in +lemma pretopology_inf : + (P ⊓ Q).pretopology (hPb.inf hQb) = P.pretopology hPb ⊓ Q.pretopology hQb := by + ext X S + exact ⟨fun hS ↦ ⟨fun hf ↦ (hS hf).left, fun hf ↦ (hS hf).right⟩, + fun h ↦ fun hf ↦ ⟨h.left hf, h.right hf⟩⟩ + +end CategoryTheory.MorphismProperty From 9800ff6dbcb799b3fe7c7cb25eb570e77893f4ee Mon Sep 17 00:00:00 2001 From: Wojciech Nawrocki Date: Thu, 17 Oct 2024 08:31:07 +0000 Subject: [PATCH 061/505] feat(LinearAlgebra/Matrix): characterize ideals in matrix rings (#17750) - Define left and two-sided ideals in a matrix ring. - Show that two-sided matrix ideals are order-isomorphic to two-sided ideals in the base ring. This code is taken from FLT, from [here](https://github.com/ImperialCollegeLondon/FLT/blob/daac491a4fa817ec39595e0a5ce2240ceed940fc/FLT/for_mathlib/CrazySimple.lean). Co-authored-by: Wojciech Nawrocki --- Mathlib.lean | 1 + Mathlib/LinearAlgebra/Matrix/Ideal.lean | 203 ++++++++++++++++++++++++ 2 files changed, 204 insertions(+) create mode 100644 Mathlib/LinearAlgebra/Matrix/Ideal.lean diff --git a/Mathlib.lean b/Mathlib.lean index d978b6190e81c..cb62c7d005329 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3089,6 +3089,7 @@ import Mathlib.LinearAlgebra.Matrix.GeneralLinearGroup.Defs import Mathlib.LinearAlgebra.Matrix.Gershgorin import Mathlib.LinearAlgebra.Matrix.Hermitian import Mathlib.LinearAlgebra.Matrix.HermitianFunctionalCalculus +import Mathlib.LinearAlgebra.Matrix.Ideal import Mathlib.LinearAlgebra.Matrix.InvariantBasisNumber import Mathlib.LinearAlgebra.Matrix.IsDiag import Mathlib.LinearAlgebra.Matrix.LDL diff --git a/Mathlib/LinearAlgebra/Matrix/Ideal.lean b/Mathlib/LinearAlgebra/Matrix/Ideal.lean new file mode 100644 index 0000000000000..2d0f2ed989b04 --- /dev/null +++ b/Mathlib/LinearAlgebra/Matrix/Ideal.lean @@ -0,0 +1,203 @@ +/- +Copyright (c) 2024 Jujian Zhang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jujian Zhang, Wojciech Nawrocki +-/ +import Mathlib.Data.Matrix.Basis +import Mathlib.RingTheory.TwoSidedIdeal.Operations + +/-! +# Ideals in a matrix ring + +This file defines left (resp. two-sided) ideals in a matrix semiring (resp. ring) +over left (resp. two-sided) ideals in the base semiring (resp. ring). + +## Main results + +* `TwoSidedIdeal.equivMatricesOver` and `TwoSidedIdeal.orderIsoMatricesOver` + establish an order isomorphism between two-sided ideals in $R$ and those in $Mₙ(R)$. +-/ + +/-! ### Left ideals in a matrix ring -/ + +namespace Ideal +open Matrix + +variable {R : Type*} [Semiring R] + (n : Type*) [Fintype n] [DecidableEq n] + +/-- The left ideal of matrices with entries in `I ≤ R`. -/ +def matricesOver (I : Ideal R) : Ideal (Matrix n n R) where + carrier := { M | ∀ i j, M i j ∈ I } + add_mem' ha hb i j := I.add_mem (ha i j) (hb i j) + zero_mem' _ _ := I.zero_mem + smul_mem' M N hN := by + intro i j + rw [smul_eq_mul, mul_apply] + apply sum_mem + intro k _ + apply I.mul_mem_left _ (hN k j) + +@[simp] +theorem mem_matricesOver (I : Ideal R) (M : Matrix n n R) : + M ∈ I.matricesOver n ↔ ∀ i j, M i j ∈ I := by rfl + +theorem matricesOver_monotone : Monotone (matricesOver (R := R) n) := + fun _ _ IJ _ MI i j => IJ (MI i j) + +theorem matricesOver_strictMono_of_nonempty [Nonempty n] : + StrictMono (matricesOver (R := R) n) := + matricesOver_monotone n |>.strictMono_of_injective <| fun I J eq => by + ext x + have : (∀ _ _, x ∈ I) ↔ (∀ _ _, x ∈ J) := congr((Matrix.of fun _ _ => x) ∈ $eq) + simpa only [forall_const] using this + +@[simp] +theorem matricesOver_bot : (⊥ : Ideal R).matricesOver n = ⊥ := by + ext M + simp only [mem_matricesOver, mem_bot] + constructor + · intro H; ext; apply H + · intro H; simp [H] + +@[simp] +theorem matricesOver_top : (⊤ : Ideal R).matricesOver n = ⊤ := by + ext; simp + +end Ideal + +/-! ### Two-sided ideals in a matrix ring -/ + +namespace TwoSidedIdeal +open Matrix + +variable {R : Type*} [Ring R] + (n : Type*) [Fintype n] + +/-- The two-sided ideal of matrices with entries in `I ≤ R`. -/ +def matricesOver (I : TwoSidedIdeal R) : TwoSidedIdeal (Matrix n n R) := + TwoSidedIdeal.mk' { M | ∀ i j, M i j ∈ I } + (fun _ _ => I.zero_mem) + (fun ha hb i j => I.add_mem (ha i j) (hb i j)) + (fun ha i j => I.neg_mem (ha i j)) + (fun ha i j => by + rw [mul_apply] + apply sum_mem + intro k _ + apply I.mul_mem_left _ _ (ha k j)) + (fun ha i j => by + rw [mul_apply] + apply sum_mem + intro k _ + apply I.mul_mem_right _ _ (ha i k)) + +@[simp] +lemma mem_matricesOver (I : TwoSidedIdeal R) (M : Matrix n n R) : + M ∈ I.matricesOver n ↔ ∀ i j, M i j ∈ I := by + simp [matricesOver] + +theorem matricesOver_monotone : Monotone (matricesOver (R := R) n) := + fun _ _ IJ _ MI i j => IJ (MI i j) + +theorem matricesOver_strictMono_of_nonempty [h : Nonempty n] : + StrictMono (matricesOver (R := R) n) := + matricesOver_monotone n |>.strictMono_of_injective <| fun I J eq => by + ext x + have : _ ↔ _ := congr((Matrix.of fun _ _ => x) ∈ $eq) + simpa only [mem_matricesOver, of_apply, forall_const] using this + +@[simp] +theorem matricesOver_bot : (⊥ : TwoSidedIdeal R).matricesOver n = ⊥ := by + ext M + simp only [mem_matricesOver, mem_bot] + constructor + · intro H; ext; apply H + · intro H; simp [H] + +@[simp] +theorem matricesOver_top : (⊤ : TwoSidedIdeal R).matricesOver n = ⊤ := by + ext; simp + +theorem asIdeal_matricesOver [DecidableEq n] (I : TwoSidedIdeal R) : + asIdeal (I.matricesOver n) = (asIdeal I).matricesOver n := by + ext; simp + +variable {n : Type*} [Fintype n] [DecidableEq n] + +/-- +Two-sided ideals in $R$ correspond bijectively to those in $Mₙ(R)$. +Given an ideal $I ≤ R$, we send it to $Mₙ(I)$. +Given an ideal $J ≤ Mₙ(R)$, we send it to $\{Nᵢⱼ ∣ ∃ N ∈ J\}$. +-/ +@[simps] +def equivMatricesOver (i j : n) : TwoSidedIdeal R ≃ TwoSidedIdeal (Matrix n n R) where + toFun I := I.matricesOver n + invFun J := TwoSidedIdeal.mk' + { N i j | N ∈ J } + ⟨0, J.zero_mem, rfl⟩ + (by rintro _ _ ⟨x, hx, rfl⟩ ⟨y, hy, rfl⟩; exact ⟨x + y, J.add_mem hx hy, rfl⟩) + (by rintro _ ⟨x, hx, rfl⟩; exact ⟨-x, J.neg_mem hx, rfl⟩) + (by + rintro x _ ⟨y, hy, rfl⟩ + exact ⟨diagonal (fun _ ↦ x) * y, J.mul_mem_left _ _ hy, by simp⟩) + (by + rintro _ y ⟨x, hx, rfl⟩ + exact ⟨x * diagonal (fun _ ↦ y), J.mul_mem_right _ _ hx, by simp⟩) + right_inv J := SetLike.ext fun x ↦ by + simp only [mem_mk', Set.mem_image, SetLike.mem_coe, mem_matricesOver] + constructor + · intro h + choose y hy1 hy2 using h + rw [matrix_eq_sum_stdBasisMatrix x] + refine sum_mem fun k _ ↦ sum_mem fun l _ ↦ ?_ + suffices + stdBasisMatrix k l (x k l) = + stdBasisMatrix k i 1 * y k l * stdBasisMatrix j l 1 by + rw [this] + exact J.mul_mem_right _ _ (J.mul_mem_left _ _ <| hy1 _ _) + ext a b + by_cases hab : a = k ∧ b = l + · rcases hab with ⟨ha, hb⟩ + subst ha hb + simp only [stdBasisMatrix, and_self, ↓reduceIte, StdBasisMatrix.mul_right_apply_same, + StdBasisMatrix.mul_left_apply_same, one_mul, mul_one] + rw [hy2 a b] + · conv_lhs => + dsimp [stdBasisMatrix] + rw [if_neg (by tauto)] + rw [not_and_or] at hab + rcases hab with ha | hb + · rw [mul_assoc, StdBasisMatrix.mul_left_apply_of_ne (h := ha)] + · rw [StdBasisMatrix.mul_right_apply_of_ne (hbj := hb)] + · intro hx k l + refine ⟨stdBasisMatrix i k 1 * x * stdBasisMatrix l j 1, + J.mul_mem_right _ _ (J.mul_mem_left _ _ hx), ?_⟩ + rw [StdBasisMatrix.mul_right_apply_same, StdBasisMatrix.mul_left_apply_same, + mul_one, one_mul] + left_inv I := SetLike.ext fun x ↦ by + simp only [mem_mk', Set.mem_image, SetLike.mem_coe, mem_matricesOver] + constructor + · intro h + choose y hy1 hy2 using h + exact hy2 ▸ hy1 _ _ + · intro h + exact ⟨of fun _ _ => x, by simp [h], rfl⟩ + +/-- +Two-sided ideals in $R$ are order-isomorphic with those in $Mₙ(R)$. +See also `equivMatricesOver`. +-/ +@[simps!] +def orderIsoMatricesOver (i j : n) : TwoSidedIdeal R ≃o TwoSidedIdeal (Matrix n n R) where + __ := equivMatricesOver i j + map_rel_iff' {I J} := by + simp only [equivMatricesOver_apply] + constructor + · intro le x xI + specialize @le (of fun _ _ => x) (by simp [xI]) + letI : Inhabited n := ⟨i⟩ + simpa using le + · intro IJ M MI i j + exact IJ <| MI i j + +end TwoSidedIdeal From ffb2f0dbf80ad18fa75f85bd0facee6737b59acf Mon Sep 17 00:00:00 2001 From: Riccardo Brasca Date: Thu, 17 Oct 2024 08:31:08 +0000 Subject: [PATCH 062/505] feat: add `Module.finrank_quotient_add_finrank_le` (#17782) --- Mathlib/LinearAlgebra/Dimension/Finite.lean | 9 +++++++++ Mathlib/LinearAlgebra/Dimension/StrongRankCondition.lean | 7 +++++++ 2 files changed, 16 insertions(+) diff --git a/Mathlib/LinearAlgebra/Dimension/Finite.lean b/Mathlib/LinearAlgebra/Dimension/Finite.lean index 90c82db1cfcc3..bf110e9d64974 100644 --- a/Mathlib/LinearAlgebra/Dimension/Finite.lean +++ b/Mathlib/LinearAlgebra/Dimension/Finite.lean @@ -7,6 +7,7 @@ import Mathlib.Algebra.Module.Torsion import Mathlib.SetTheory.Cardinal.Cofinality import Mathlib.LinearAlgebra.FreeModule.Finite.Basic import Mathlib.LinearAlgebra.Dimension.StrongRankCondition +import Mathlib.LinearAlgebra.Dimension.Constructions /-! # Conditions for rank to be finite @@ -426,6 +427,14 @@ theorem Module.finrank_zero_iff [NoZeroSMulDivisors R M] : rw [← rank_zero_iff (R := R), ← finrank_eq_rank] norm_cast +/-- Similar to `rank_quotient_add_rank_le` but for `finrank` and a finite `M`. -/ +lemma Module.finrank_quotient_add_finrank_le (N : Submodule R M) : + finrank R (M ⧸ N) + finrank R N ≤ finrank R M := by + haveI := nontrivial_of_invariantBasisNumber R + have := rank_quotient_add_rank_le N + rw [← finrank_eq_rank R M, ← finrank_eq_rank R, ← N.finrank_eq_rank] at this + exact mod_cast this + end StrongRankCondition theorem Module.finrank_eq_zero_of_rank_eq_zero (h : Module.rank R M = 0) : diff --git a/Mathlib/LinearAlgebra/Dimension/StrongRankCondition.lean b/Mathlib/LinearAlgebra/Dimension/StrongRankCondition.lean index 26dc00965d15b..7b753996882f1 100644 --- a/Mathlib/LinearAlgebra/Dimension/StrongRankCondition.lean +++ b/Mathlib/LinearAlgebra/Dimension/StrongRankCondition.lean @@ -443,6 +443,13 @@ noncomputable instance {R M : Type*} [DivisionRing R] [AddCommGroup M] [Module R theorem finrank_eq_rank [Module.Finite R M] : ↑(finrank R M) = Module.rank R M := by rw [Module.finrank, cast_toNat_of_lt_aleph0 (rank_lt_aleph0 R M)] +/-- If `M` is finite, then `finrank N = rank N` for all `N : Submodule M`. Note that +such an `N` need not be finitely generated. -/ +protected theorem _root_.Submodule.finrank_eq_rank [Module.Finite R M] (N : Submodule R M) : + finrank R N = Module.rank R N := by + rw [finrank, Cardinal.cast_toNat_of_lt_aleph0] + exact lt_of_le_of_lt (Submodule.rank_le N) (rank_lt_aleph0 R M) + end Module open Module From 7b3af7ebd252c921171d5d635a03288f91264ba5 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Thu, 17 Oct 2024 08:51:30 +0000 Subject: [PATCH 063/505] chore: adjust set_range_forget_map_inclusion' to the statement that is really needed (#17278) --- Mathlib/Topology/Category/TopCat/Opens.lean | 10 +++++----- scripts/no_lints_prime_decls.txt | 1 + 2 files changed, 6 insertions(+), 5 deletions(-) diff --git a/Mathlib/Topology/Category/TopCat/Opens.lean b/Mathlib/Topology/Category/TopCat/Opens.lean index 41c2df48911b0..eb8eb48fb88be 100644 --- a/Mathlib/Topology/Category/TopCat/Opens.lean +++ b/Mathlib/Topology/Category/TopCat/Opens.lean @@ -335,22 +335,22 @@ theorem functor_obj_map_obj {X Y : TopCat} {f : X ⟶ Y} (hf : IsOpenMap f) (U : exact ⟨x, hx, rfl⟩ -- Porting note: added to ease the proof of `functor_map_eq_inf` -lemma set_range_forget_map_inclusion' {X : TopCat} (U : Opens X) : - Set.range ((forget TopCat).map (inclusion' U)) = (U : Set X) := by +lemma set_range_inclusion' {X : TopCat} (U : Opens X) : + Set.range (inclusion' U) = (U : Set X) := by ext x constructor · rintro ⟨x, rfl⟩ exact x.2 · intro h exact ⟨⟨x, h⟩, rfl⟩ +@[deprecated (since := "2024-09-07")] alias set_range_forget_map_inclusion' := set_range_inclusion' @[simp] theorem functor_map_eq_inf {X : TopCat} (U V : Opens X) : U.openEmbedding.isOpenMap.functor.obj ((Opens.map U.inclusion').obj V) = V ⊓ U := by ext1 - refine Set.image_preimage_eq_inter_range.trans ?_ - erw [set_range_forget_map_inclusion' U] - rfl + simp only [IsOpenMap.functor_obj_coe, map_coe, coe_inf, + Set.image_preimage_eq_inter_range, set_range_inclusion' U] theorem map_functor_eq' {X U : TopCat} (f : U ⟶ X) (hf : OpenEmbedding f) (V) : ((Opens.map f).obj <| hf.isOpenMap.functor.obj V) = V := diff --git a/scripts/no_lints_prime_decls.txt b/scripts/no_lints_prime_decls.txt index b2fa718c5be98..8affabbf01733 100644 --- a/scripts/no_lints_prime_decls.txt +++ b/scripts/no_lints_prime_decls.txt @@ -4652,6 +4652,7 @@ TopologicalSpace.Opens.map_functor_eq' TopologicalSpace.Opens.map_id_obj' TopologicalSpace.Opens.openEmbedding' TopologicalSpace.Opens.set_range_forget_map_inclusion' +TopologicalSpace.Opens.set_range_inclusion' TopologicalSpace.SecondCountableTopology.mk' Topology.WithScott.isOpen_iff_isUpperSet_and_scottHausdorff_open' top_sdiff' From c1d3ebd51a1725828121397ad4a8f36aba2f3a07 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Thu, 17 Oct 2024 08:51:31 +0000 Subject: [PATCH 064/505] chore: clean up "was simp" porting notes (#17827) Go through all the files with porting notes along the lines of "was `simp`", "was `simp_rw`", "used to be `simp`" and try to: * Restore original, if possible; * Use another cleaner way to solve the issue; * Delete the note if it's not worth going back to the old version; * Renumber the note if it's a more applicable issue (#11119 in particular); * Or not touch the note if it's still relevant. Co-authored-by: grunweg --- .../Algebra/Category/ModuleCat/Adjunctions.lean | 11 +++-------- Mathlib/Algebra/Module/LinearMap/Prod.lean | 4 +--- Mathlib/Algebra/Polynomial/Coeff.lean | 1 - Mathlib/Analysis/Calculus/Implicit.lean | 6 +----- Mathlib/Analysis/Normed/Group/AddTorsor.lean | 6 ++---- Mathlib/Analysis/Normed/Lp/PiLp.lean | 1 - Mathlib/Analysis/RCLike/Basic.lean | 12 ++++++------ Mathlib/Data/Finsupp/AList.lean | 3 +-- Mathlib/Data/Seq/WSeq.lean | 2 -- .../Geometry/Euclidean/Angle/Oriented/Basic.lean | 15 ++------------- Mathlib/Geometry/Euclidean/MongePoint.lean | 12 ++---------- Mathlib/GroupTheory/Perm/Fin.lean | 2 +- .../LinearAlgebra/AffineSpace/Combination.lean | 8 ++------ Mathlib/LinearAlgebra/Alternating/Basic.lean | 4 +--- .../LinearAlgebra/CliffordAlgebra/Grading.lean | 3 +-- Mathlib/LinearAlgebra/Determinant.lean | 4 +--- Mathlib/LinearAlgebra/Isomorphisms.lean | 4 +--- .../Constructions/Prod/Integral.lean | 2 -- Mathlib/MeasureTheory/Group/Integral.lean | 9 +++------ .../NumberTheory/Cyclotomic/PrimitiveRoots.lean | 9 +++------ Mathlib/RingTheory/DedekindDomain/PID.lean | 1 - Mathlib/RingTheory/DedekindDomain/SInteger.lean | 4 +--- Mathlib/SetTheory/Cardinal/Basic.lean | 2 +- Mathlib/Topology/Algebra/Module/Basic.lean | 6 +----- Mathlib/Topology/Instances/AddCircle.lean | 4 +--- 25 files changed, 35 insertions(+), 100 deletions(-) diff --git a/Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean b/Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean index 2abebe30399c0..c3f9a38ef087e 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean @@ -317,9 +317,7 @@ def lift (F : C ⥤ D) : Free R C ⥤ D where map_id := by dsimp [CategoryTheory.categoryFree]; simp map_comp {X Y Z} f g := by apply Finsupp.induction_linear f - · -- Porting note (#10959): simp used to be able to close this goal - dsimp - rw [Limits.zero_comp, sum_zero_index, Limits.zero_comp] + · simp · intro f₁ f₂ w₁ w₂ rw [add_comp] dsimp at * @@ -331,9 +329,7 @@ def lift (F : C ⥤ D) : Free R C ⥤ D where · intros; simp only [add_smul] · intro f' r apply Finsupp.induction_linear g - · -- Porting note (#10959): simp used to be able to close this goal - dsimp - rw [Limits.comp_zero, sum_zero_index, Limits.comp_zero] + · simp · intro f₁ f₂ w₁ w₂ rw [comp_add] dsimp at * @@ -376,8 +372,7 @@ def ext {F G : Free R C ⥤ D} [F.Additive] [F.Linear R] [G.Additive] [G.Linear (by intro X Y f apply Finsupp.induction_linear f - · -- Porting note (#10959): simp used to be able to close this goal - rw [Functor.map_zero, Limits.zero_comp, Functor.map_zero, Limits.comp_zero] + · simp · intro f₁ f₂ w₁ w₂ -- Porting note: Using rw instead of simp rw [Functor.map_add, add_comp, w₁, w₂, Functor.map_add, comp_add] diff --git a/Mathlib/Algebra/Module/LinearMap/Prod.lean b/Mathlib/Algebra/Module/LinearMap/Prod.lean index c21bf6e0c5a84..5b80a635be778 100644 --- a/Mathlib/Algebra/Module/LinearMap/Prod.lean +++ b/Mathlib/Algebra/Module/LinearMap/Prod.lean @@ -35,9 +35,7 @@ theorem isLinearMap_sub [AddCommGroup M] [Module R M] : IsLinearMap R fun x : M × M => x.1 - x.2 := by apply IsLinearMap.mk · intro x y - -- porting note (#10745): was `simp [add_comm, add_left_comm, sub_eq_add_neg]` - rw [Prod.fst_add, Prod.snd_add] - abel + simp [add_comm, add_assoc, add_left_comm, sub_eq_add_neg] · intro x y simp [smul_sub] diff --git a/Mathlib/Algebra/Polynomial/Coeff.lean b/Mathlib/Algebra/Polynomial/Coeff.lean index a0550adae74d1..3d6efd887e11e 100644 --- a/Mathlib/Algebra/Polynomial/Coeff.lean +++ b/Mathlib/Algebra/Polynomial/Coeff.lean @@ -104,7 +104,6 @@ lemma coeff_list_sum_map {ι : Type*} (l : List ι) (f : ι → R[X]) (n : ℕ) theorem coeff_sum [Semiring S] (n : ℕ) (f : ℕ → R → S[X]) : coeff (p.sum f) n = p.sum fun a b => coeff (f a b) n := by rcases p with ⟨⟩ - -- porting note (#10745): was `simp [Polynomial.sum, support, coeff]`. simp [Polynomial.sum, support_ofFinsupp, coeff_ofFinsupp] /-- Decomposes the coefficient of the product `p * q` as a sum diff --git a/Mathlib/Analysis/Calculus/Implicit.lean b/Mathlib/Analysis/Calculus/Implicit.lean index 2e323d8872a82..01a204f86c4db 100644 --- a/Mathlib/Analysis/Calculus/Implicit.lean +++ b/Mathlib/Analysis/Calculus/Implicit.lean @@ -190,11 +190,7 @@ theorem implicitFunction_hasStrictFDerivAt (g'inv : G →L[𝕜] E) convert this.comp (φ.rightFun φ.pt) ((hasStrictFDerivAt_const _ _).prod (hasStrictFDerivAt_id _)) -- Porting note: added parentheses to help `simp` simp only [ContinuousLinearMap.ext_iff, (ContinuousLinearMap.comp_apply)] at hg'inv hg'invf ⊢ - -- porting note (#10745): was `simp [ContinuousLinearEquiv.eq_symm_apply]`; - -- both `simp` and `rw` fail here, `erw` works - intro x - erw [ContinuousLinearEquiv.eq_symm_apply] - simp [*] + simp [ContinuousLinearEquiv.eq_symm_apply, *] end ImplicitFunctionData diff --git a/Mathlib/Analysis/Normed/Group/AddTorsor.lean b/Mathlib/Analysis/Normed/Group/AddTorsor.lean index c4b97754790bb..e3a3f8109709e 100644 --- a/Mathlib/Analysis/Normed/Group/AddTorsor.lean +++ b/Mathlib/Analysis/Normed/Group/AddTorsor.lean @@ -44,8 +44,7 @@ variable {α V P W Q : Type*} [SeminormedAddCommGroup V] [PseudoMetricSpace P] [ instance (priority := 100) NormedAddTorsor.to_isometricVAdd : IsometricVAdd V P := ⟨fun c => Isometry.of_dist_eq fun x y => by - -- porting note (#10745): was `simp [NormedAddTorsor.dist_eq_norm']` - rw [NormedAddTorsor.dist_eq_norm', NormedAddTorsor.dist_eq_norm', vadd_vsub_vadd_cancel_left]⟩ + simp [NormedAddTorsor.dist_eq_norm']⟩ /-- A `SeminormedAddCommGroup` is a `NormedAddTorsor` over itself. -/ instance (priority := 100) SeminormedAddCommGroup.toNormedAddTorsor : NormedAddTorsor V V where @@ -99,8 +98,7 @@ theorem nndist_vadd_cancel_right (v₁ v₂ : V) (x : P) : nndist (v₁ +ᵥ x) @[simp] theorem dist_vadd_left (v : V) (x : P) : dist (v +ᵥ x) x = ‖v‖ := by - -- porting note (#10745): was `simp [dist_eq_norm_vsub V _ x]` - rw [dist_eq_norm_vsub V _ x, vadd_vsub] + simp [dist_eq_norm_vsub V _ x] @[simp] theorem nndist_vadd_left (v : V) (x : P) : nndist (v +ᵥ x) x = ‖v‖₊ := diff --git a/Mathlib/Analysis/Normed/Lp/PiLp.lean b/Mathlib/Analysis/Normed/Lp/PiLp.lean index 2fed648aba090..3902519c5b489 100644 --- a/Mathlib/Analysis/Normed/Lp/PiLp.lean +++ b/Mathlib/Analysis/Normed/Lp/PiLp.lean @@ -825,7 +825,6 @@ theorem edist_equiv_symm_single_same (i : ι) (b₁ b₂ : β i) : ((WithLp.equiv p (∀ i, β i)).symm (Pi.single i b₁)) ((WithLp.equiv p (∀ i, β i)).symm (Pi.single i b₂)) = edist b₁ b₂ := by - -- Porting note: was `simpa using` simp only [edist_nndist, nndist_equiv_symm_single_same p β i b₁ b₂] end Single diff --git a/Mathlib/Analysis/RCLike/Basic.lean b/Mathlib/Analysis/RCLike/Basic.lean index 076e45a1a3a71..7db8b7b8d5973 100644 --- a/Mathlib/Analysis/RCLike/Basic.lean +++ b/Mathlib/Analysis/RCLike/Basic.lean @@ -262,7 +262,7 @@ theorem I_im (z : K) : im z * im (I : K) = im z := @[simp, rclike_simps] theorem I_im' (z : K) : im (I : K) * im z = im z := by rw [mul_comm, I_im] -@[rclike_simps] -- porting note (#10618): was `simp` +@[rclike_simps] -- porting note (#11119): was `simp` theorem I_mul_re (z : K) : re (I * z) = -im z := by simp only [I_re, zero_sub, I_im', zero_mul, mul_re] @@ -394,7 +394,7 @@ theorem normSq_one : normSq (1 : K) = 1 := theorem normSq_nonneg (z : K) : 0 ≤ normSq z := add_nonneg (mul_self_nonneg _) (mul_self_nonneg _) -@[rclike_simps] -- porting note (#10618): was `simp` +@[rclike_simps] -- porting note (#11119): was `simp` theorem normSq_eq_zero {z : K} : normSq z = 0 ↔ z = 0 := map_eq_zero _ @@ -409,7 +409,7 @@ theorem normSq_neg (z : K) : normSq (-z) = normSq z := by simp only [normSq_eq_d theorem normSq_conj (z : K) : normSq (conj z) = normSq z := by simp only [normSq_apply, neg_mul, mul_neg, neg_neg, rclike_simps] -@[rclike_simps] -- porting note (#10618): was `simp` +@[rclike_simps] -- porting note (#11119): was `simp` theorem normSq_mul (z w : K) : normSq (z * w) = normSq z * normSq w := map_mul _ z w @@ -466,7 +466,7 @@ theorem div_im (z w : K) : im (z / w) = im z * re w / normSq w - re z * im w / n simp only [div_eq_mul_inv, mul_assoc, sub_eq_add_neg, add_comm, neg_mul, mul_neg, map_neg, rclike_simps] -@[rclike_simps] -- porting note (#10618): was `simp` +@[rclike_simps] -- porting note (#11119): was `simp` theorem conj_inv (x : K) : conj x⁻¹ = (conj x)⁻¹ := star_inv' _ @@ -506,11 +506,11 @@ theorem inv_I : (I : K)⁻¹ = -I := by @[simp, rclike_simps] theorem div_I (z : K) : z / I = -(z * I) := by rw [div_eq_mul_inv, inv_I, mul_neg] -@[rclike_simps] -- porting note (#10618): was `simp` +@[rclike_simps] -- porting note (#11119): was `simp` theorem normSq_inv (z : K) : normSq z⁻¹ = (normSq z)⁻¹ := map_inv₀ normSq z -@[rclike_simps] -- porting note (#10618): was `simp` +@[rclike_simps] -- porting note (#11119): was `simp` theorem normSq_div (z w : K) : normSq (z / w) = normSq z / normSq w := map_div₀ normSq z w diff --git a/Mathlib/Data/Finsupp/AList.lean b/Mathlib/Data/Finsupp/AList.lean index d7f3418427f67..4946a43826403 100644 --- a/Mathlib/Data/Finsupp/AList.lean +++ b/Mathlib/Data/Finsupp/AList.lean @@ -103,8 +103,7 @@ theorem insert_lookupFinsupp [DecidableEq α] (l : AList fun _x : α => M) (a : theorem singleton_lookupFinsupp (a : α) (m : M) : (singleton a m).lookupFinsupp = Finsupp.single a m := by classical - -- porting note (#10745): was `simp [← AList.insert_empty]` but timeout issues - simp only [← AList.insert_empty, insert_lookupFinsupp, empty_lookupFinsupp, Finsupp.zero_update] + simp [← AList.insert_empty] @[simp] theorem _root_.Finsupp.toAList_lookupFinsupp (f : α →₀ M) : f.toAList.lookupFinsupp = f := by diff --git a/Mathlib/Data/Seq/WSeq.lean b/Mathlib/Data/Seq/WSeq.lean index d6bd6ad6701ca..5b57888ef3f6e 100644 --- a/Mathlib/Data/Seq/WSeq.lean +++ b/Mathlib/Data/Seq/WSeq.lean @@ -627,7 +627,6 @@ theorem dropn_cons (a : α) (s) (n) : drop (cons a s) (n + 1) = drop s n := by induction n with | zero => simp [drop] | succ n n_ih => - -- porting note (#10745): was `simp [*, drop]`. simp [drop, ← n_ih] @[simp] @@ -875,7 +874,6 @@ theorem exists_get?_of_mem {s : WSeq α} {a} (h : a ∈ s) : ∃ n, some a ∈ g apply ret_mem · cases' h with n h exists n + 1 - -- porting note (#10745): was `simp [get?]`. simpa [get?] · intro s' h cases' h with n h diff --git a/Mathlib/Geometry/Euclidean/Angle/Oriented/Basic.lean b/Mathlib/Geometry/Euclidean/Angle/Oriented/Basic.lean index 3f2938cdce408..0ad26516219b6 100644 --- a/Mathlib/Geometry/Euclidean/Angle/Oriented/Basic.lean +++ b/Mathlib/Geometry/Euclidean/Angle/Oriented/Basic.lean @@ -757,19 +757,8 @@ outside of that proof. -/ theorem oangle_smul_add_right_eq_zero_or_eq_pi_iff {x y : V} (r : ℝ) : o.oangle x (r • x + y) = 0 ∨ o.oangle x (r • x + y) = π ↔ o.oangle x y = 0 ∨ o.oangle x y = π := by - simp_rw [oangle_eq_zero_or_eq_pi_iff_not_linearIndependent, Fintype.not_linearIndependent_iff] - -- Porting note: at this point all occurrences of the bound variable `i` are of type - -- `Fin (Nat.succ (Nat.succ 0))`, but `Fin.sum_univ_two` and `Fin.exists_fin_two` expect it to be - -- `Fin 2` instead. Hence all the `conv`s. - -- Was `simp_rw [Fin.sum_univ_two, Fin.exists_fin_two]` - conv_lhs => enter [1, g, 1, 1, 2, i]; tactic => change Fin 2 at i - conv_lhs => enter [1, g]; rw [Fin.sum_univ_two] - conv_rhs => enter [1, g, 1, 1, 2, i]; tactic => change Fin 2 at i - conv_rhs => enter [1, g]; rw [Fin.sum_univ_two] - conv_lhs => enter [1, g, 2, 1, i]; tactic => change Fin 2 at i - conv_lhs => enter [1, g]; rw [Fin.exists_fin_two] - conv_rhs => enter [1, g, 2, 1, i]; tactic => change Fin 2 at i - conv_rhs => enter [1, g]; rw [Fin.exists_fin_two] + simp_rw [oangle_eq_zero_or_eq_pi_iff_not_linearIndependent, Fintype.not_linearIndependent_iff, + Fin.sum_univ_two, Fin.exists_fin_two] refine ⟨fun h => ?_, fun h => ?_⟩ · rcases h with ⟨m, h, hm⟩ change m 0 • x + m 1 • (r • x + y) = 0 at h diff --git a/Mathlib/Geometry/Euclidean/MongePoint.lean b/Mathlib/Geometry/Euclidean/MongePoint.lean index fbf488bef44b3..d829d76dfb4ef 100644 --- a/Mathlib/Geometry/Euclidean/MongePoint.lean +++ b/Mathlib/Geometry/Euclidean/MongePoint.lean @@ -502,16 +502,8 @@ theorem dist_orthocenter_reflection_circumcenter (t : Triangle ℝ P) {i₁ i₂ ∃ i₃, univ \ ({i₁, i₂} : Finset (Fin 3)) = {i₃} ∧ i₃ ≠ i₁ ∧ i₃ ≠ i₂ := by -- Porting note (#11043): was `decide!` fin_cases i₁ <;> fin_cases i₂ <;> simp at h <;> decide - -- Porting note: Original proof was `simp_rw [← sum_sdiff hu, hi₃]; simp [hi₃₁, hi₃₂]; norm_num` - rw [← sum_sdiff hu, ← sum_sdiff hu, hi₃, sum_singleton, ← sum_sdiff hu, hi₃] - split_ifs with h - · exact (h.elim hi₃₁ hi₃₂).elim - simp only [zero_add, Nat.cast_one, inv_one, sub_zero, one_mul, pointsWithCircumcenter_point, - sum_singleton, h, ite_false, dist_self, mul_zero, mem_singleton, true_or, ite_true, sub_self, - zero_mul, implies_true, sum_insert_of_eq_zero_if_not_mem, or_true, add_zero, div_one, - sub_neg_eq_add, pointsWithCircumcenter_eq_circumcenter, dist_circumcenter_eq_circumradius, - sum_const_zero, dist_circumcenter_eq_circumradius', mul_one, neg_add_rev, add_self_div_two] - norm_num + simp_rw [← sum_sdiff hu, hi₃] + norm_num [hi₃₁, hi₃₂] /-- The distance from the orthocenter to the reflection of the circumcenter in a side equals the circumradius, variant using a diff --git a/Mathlib/GroupTheory/Perm/Fin.lean b/Mathlib/GroupTheory/Perm/Fin.lean index 7df1d1ceaefbe..6d0fa972e68ea 100644 --- a/Mathlib/GroupTheory/Perm/Fin.lean +++ b/Mathlib/GroupTheory/Perm/Fin.lean @@ -151,7 +151,7 @@ theorem cycleRange_of_gt {n : ℕ} {i j : Fin n.succ} (h : i < j) : cycleRange i theorem cycleRange_of_le {n : ℕ} {i j : Fin n.succ} (h : j ≤ i) : cycleRange i j = if j = i then 0 else j + 1 := by cases n - · exact Subsingleton.elim (α := Fin 1) _ _ --Porting note; was `simp` + · subsingleton have : j = (Fin.castLE (Nat.succ_le_of_lt i.is_lt)) ⟨j, lt_of_le_of_lt h (Nat.lt_succ_self i)⟩ := by simp ext diff --git a/Mathlib/LinearAlgebra/AffineSpace/Combination.lean b/Mathlib/LinearAlgebra/AffineSpace/Combination.lean index 8d58836617812..c3adfc82235fd 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/Combination.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/Combination.lean @@ -935,10 +935,7 @@ theorem affineCombination_mem_affineSpan [Nontrivial k] {s : Finset ι} {w : ι have hv : s.affineCombination k p w -ᵥ p i1 ∈ (affineSpan k (Set.range p)).direction := by rw [direction_affineSpan, ← hw1s, Finset.affineCombination_vsub] apply weightedVSub_mem_vectorSpan - -- Porting note: Rest was `simp [Pi.sub_apply, h, hw1]`, - -- but `Pi.sub_apply` transforms the goal into nonsense - change (Finset.sum s fun i => w i - w1 i) = 0 - simp only [Finset.sum_sub_distrib, h, hw1, sub_self] + simp [Pi.sub_apply, h, hw1] rw [← vsub_vadd (s.affineCombination k p w) (p i1)] exact AffineSubspace.vadd_mem_of_mem_direction hv (mem_affineSpan k (Set.mem_range_self _)) @@ -1019,8 +1016,7 @@ theorem eq_affineCombination_of_mem_affineSpan {p1 : P} {p : ι → P} s'.affineCombination_of_eq_one_of_eq_zero w0 p (Finset.mem_insert_self _ _) (Function.update_same _ _ _) fun _ _ hne => Function.update_noteq hne _ _ refine ⟨s', w0 + w', ?_, ?_⟩ - · -- Porting note: proof was `simp [Pi.add_apply, Finset.sum_add_distrib, hw0, h']` - simp only [Pi.add_apply, Finset.sum_add_distrib, hw0, h', add_zero] + · simp [Pi.add_apply, Finset.sum_add_distrib, hw0, h'] · rw [add_comm, ← Finset.weightedVSub_vadd_affineCombination, hw0s, hs', vsub_vadd] theorem eq_affineCombination_of_mem_affineSpan_of_fintype [Fintype ι] {p1 : P} {p : ι → P} diff --git a/Mathlib/LinearAlgebra/Alternating/Basic.lean b/Mathlib/LinearAlgebra/Alternating/Basic.lean index 69536a10962b0..994b3b3d983db 100644 --- a/Mathlib/LinearAlgebra/Alternating/Basic.lean +++ b/Mathlib/LinearAlgebra/Alternating/Basic.lean @@ -491,9 +491,7 @@ theorem add_compLinearMap (f₁ f₂ : M [⋀^ι]→ₗ[R] N) (g : M₂ →ₗ[R theorem compLinearMap_zero [Nonempty ι] (f : M [⋀^ι]→ₗ[R] N) : f.compLinearMap (0 : M₂ →ₗ[R] M) = 0 := by ext - -- Porting note: Was `simp_rw [.., ← Pi.zero_def, map_zero, zero_apply]`. - simp_rw [compLinearMap_apply, LinearMap.zero_apply, ← Pi.zero_def, zero_apply] - exact map_zero f + simp_rw [compLinearMap_apply, LinearMap.zero_apply, ← Pi.zero_def, map_zero, zero_apply] /-- Composing an alternating map with the identity linear map in each argument. -/ @[simp] diff --git a/Mathlib/LinearAlgebra/CliffordAlgebra/Grading.lean b/Mathlib/LinearAlgebra/CliffordAlgebra/Grading.lean index e5f6c2be85383..145bb1e7b4a15 100644 --- a/Mathlib/LinearAlgebra/CliffordAlgebra/Grading.lean +++ b/Mathlib/LinearAlgebra/CliffordAlgebra/Grading.lean @@ -214,8 +214,7 @@ theorem odd_induction {P : ∀ x, x ∈ evenOdd Q 1 → Prop} (zero_add (1 : ZMod 2) ▸ SetLike.mul_mem_graded (ι_mul_ι_mem_evenOdd_zero Q m₁ m₂) hx)) (x : CliffordAlgebra Q) (hx : x ∈ evenOdd Q 1) : P x hx := by refine evenOdd_induction _ _ (motive := P) (fun ιv => ?_) add ι_mul_ι_mul x hx - -- Porting note: was `simp_rw [ZMod.val_one, pow_one]`, lean4#1926 - intro h; rw [ZMod.val_one, pow_one] at h; revert h + simp_rw [ZMod.val_one, pow_one] rintro ⟨v, rfl⟩ exact ι v diff --git a/Mathlib/LinearAlgebra/Determinant.lean b/Mathlib/LinearAlgebra/Determinant.lean index dc9936ab62f2a..c33431bdc1398 100644 --- a/Mathlib/LinearAlgebra/Determinant.lean +++ b/Mathlib/LinearAlgebra/Determinant.lean @@ -599,9 +599,7 @@ theorem Basis.det_unitsSMul (e : Basis ι R M) (w : ι → Rˣ) : (↑(∏ i, w i)⁻¹ : R) • Matrix.det fun i j => e.repr (f j) i simp only [e.repr_unitsSMul] convert Matrix.det_mul_column (fun i => (↑(w i)⁻¹ : R)) fun i j => e.repr (f j) i - -- porting note (#10745): was `simp [← Finset.prod_inv_distrib]` - simp only [← Finset.prod_inv_distrib] - norm_cast + simp [← Finset.prod_inv_distrib] /-- The determinant of a basis constructed by `unitsSMul` is the product of the given units. -/ @[simp] diff --git a/Mathlib/LinearAlgebra/Isomorphisms.lean b/Mathlib/LinearAlgebra/Isomorphisms.lean index 54a39b6372052..998bbedc53b52 100644 --- a/Mathlib/LinearAlgebra/Isomorphisms.lean +++ b/Mathlib/LinearAlgebra/Isomorphisms.lean @@ -120,9 +120,7 @@ theorem quotientInfEquivSupQuotient_symm_apply_left (p p' : Submodule R M) (x : -- @[simp] -- Porting note (#10618): simp can prove this theorem quotientInfEquivSupQuotient_symm_apply_eq_zero_iff {p p' : Submodule R M} {x : ↥(p ⊔ p')} : (quotientInfEquivSupQuotient p p').symm (Submodule.Quotient.mk x) = 0 ↔ (x : M) ∈ p' := - (LinearEquiv.symm_apply_eq _).trans <| by - -- porting note (#10745): was `simp`. - rw [_root_.map_zero, Quotient.mk_eq_zero, mem_comap, Submodule.coe_subtype] + (LinearEquiv.symm_apply_eq _).trans <| by simp theorem quotientInfEquivSupQuotient_symm_apply_right (p p' : Submodule R M) {x : ↥(p ⊔ p')} (hx : (x : M) ∈ p') : (quotientInfEquivSupQuotient p p').symm (Submodule.Quotient.mk x) diff --git a/Mathlib/MeasureTheory/Constructions/Prod/Integral.lean b/Mathlib/MeasureTheory/Constructions/Prod/Integral.lean index 25d335405cc0f..c86c172d3f6b7 100644 --- a/Mathlib/MeasureTheory/Constructions/Prod/Integral.lean +++ b/Mathlib/MeasureTheory/Constructions/Prod/Integral.lean @@ -153,7 +153,6 @@ theorem integrable_measure_prod_mk_left {s : Set (α × β)} (hs : MeasurableSet refine ⟨(measurable_measure_prod_mk_left hs).ennreal_toReal.aemeasurable.aestronglyMeasurable, ?_⟩ simp_rw [HasFiniteIntegral, ennnorm_eq_ofReal toReal_nonneg] convert h2s.lt_top using 1 - -- Porting note: was `simp_rw` rw [prod_apply hs] apply lintegral_congr_ae filter_upwards [ae_measure_lt_top hs h2s] with x hx @@ -426,7 +425,6 @@ theorem integral_prod (f : α × β → E) (hf : Integrable f (μ.prod ν)) : integral_indicator (measurable_prod_mk_left hs), setIntegral_const, integral_smul_const, integral_toReal (measurable_measure_prod_mk_left hs).aemeasurable (ae_measure_lt_top hs h2s.ne)] - -- Porting note: was `simp_rw` rw [prod_apply hs] · rintro f g - i_f i_g hf hg simp_rw [integral_add' i_f i_g, integral_integral_add' i_f i_g, hf, hg] diff --git a/Mathlib/MeasureTheory/Group/Integral.lean b/Mathlib/MeasureTheory/Group/Integral.lean index 665790505f3cd..31e10db1f25bb 100644 --- a/Mathlib/MeasureTheory/Group/Integral.lean +++ b/Mathlib/MeasureTheory/Group/Integral.lean @@ -68,9 +68,7 @@ theorem integral_mul_right_eq_self [IsMulRightInvariant μ] (f : G → E) (g : G @[to_additive] -- Porting note: was `@[simp]` theorem integral_div_right_eq_self [IsMulRightInvariant μ] (f : G → E) (g : G) : (∫ x, f (x / g) ∂μ) = ∫ x, f x ∂μ := by - simp_rw [div_eq_mul_inv] - -- Porting note: was `simp_rw` - rw [integral_mul_right_eq_self f g⁻¹] + simp_rw [div_eq_mul_inv, integral_mul_right_eq_self f g⁻¹] /-- If some left-translate of a function negates it, then the integral of the function with respect to a left-invariant measure is 0. -/ @@ -123,9 +121,8 @@ theorem integrable_comp_div_left (f : G → F) [IsInvInvariant μ] [IsMulLeftInv @[to_additive] -- Porting note: was `@[simp]` theorem integral_div_left_eq_self (f : G → E) (μ : Measure G) [IsInvInvariant μ] [IsMulLeftInvariant μ] (x' : G) : (∫ x, f (x' / x) ∂μ) = ∫ x, f x ∂μ := by - simp_rw [div_eq_mul_inv] - -- Porting note: was `simp_rw` - rw [integral_inv_eq_self (fun x => f (x' * x)) μ, integral_mul_left_eq_self f x'] + simp_rw [div_eq_mul_inv, integral_inv_eq_self (fun x => f (x' * x)) μ, + integral_mul_left_eq_self f x'] end MeasurableMul diff --git a/Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean b/Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean index 0d5dc764a96e2..0e8d24cf92d00 100644 --- a/Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean +++ b/Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean @@ -489,8 +489,7 @@ theorem norm_pow_sub_one_two {k : ℕ} (hζ : IsPrimitiveRoot ζ (2 ^ (k + 1))) have H : (-1 : L) - (1 : L) = algebraMap K L (-2) := by simp only [map_neg, map_ofNat] ring --- Porting note: `simpa using hirr` was `simp [hirr]`. - replace hirr : Irreducible (cyclotomic ((2 : ℕ+) ^ (k + 1) : ℕ+) K) := by simpa using hirr + replace hirr : Irreducible (cyclotomic ((2 : ℕ+) ^ (k + 1) : ℕ+) K) := by simp [hirr] -- Porting note: the proof is slightly different because of coercions. rw [this.eq_neg_one_of_two_right, H, Algebra.norm_algebraMap, IsCyclotomicExtension.finrank L hirr, pow_coe, show ((2 : ℕ+) : ℕ) = 2 from rfl, @@ -506,10 +505,8 @@ theorem norm_sub_one_two {k : ℕ} (hζ : IsPrimitiveRoot ζ (2 ^ k)) (hk : 2 simp only [← coe_lt_coe, one_coe, pow_coe] nth_rw 1 [← pow_one 2] exact pow_lt_pow_right one_lt_two (lt_of_lt_of_le one_lt_two hk) --- Porting note: `simpa using hirr` was `simp [hirr]`_ - replace hirr : Irreducible (cyclotomic ((2 : ℕ+) ^ k : ℕ+) K) := by simpa using hirr --- Porting note: `simpa using hζ` was `simp [hζ]`_ - replace hζ : IsPrimitiveRoot ζ (2 ^ k : ℕ+) := by simpa using hζ + replace hirr : Irreducible (cyclotomic ((2 : ℕ+) ^ k : ℕ+) K) := by simp [hirr] + replace hζ : IsPrimitiveRoot ζ (2 ^ k : ℕ+) := by simp [hζ] obtain ⟨k₁, hk₁⟩ := exists_eq_succ_of_ne_zero (lt_of_lt_of_le zero_lt_two hk).ne.symm -- Porting note: the proof is slightly different because of coercions. simpa [hk₁, show ((2 : ℕ+) : ℕ) = 2 from rfl] using sub_one_norm_eq_eval_cyclotomic hζ this hirr diff --git a/Mathlib/RingTheory/DedekindDomain/PID.lean b/Mathlib/RingTheory/DedekindDomain/PID.lean index 258435cca7f9a..6bc5095ed0a68 100644 --- a/Mathlib/RingTheory/DedekindDomain/PID.lean +++ b/Mathlib/RingTheory/DedekindDomain/PID.lean @@ -42,7 +42,6 @@ theorem Ideal.eq_span_singleton_of_mem_of_not_mem_sq_of_not_mem_prime_ne {P : Id exact hxP2 (zero_mem _) by_cases hP0 : P = ⊥ · subst hP0 - -- Porting note: was `simpa using hxP2` but that hypothesis didn't even seem relevant in Lean 3 rwa [eq_comm, span_singleton_eq_bot, ← mem_bot] have hspan0 : span ({x} : Set R) ≠ ⊥ := mt Ideal.span_singleton_eq_bot.mp hx0 have span_le := (Ideal.span_singleton_le_iff_mem _).mpr x_mem diff --git a/Mathlib/RingTheory/DedekindDomain/SInteger.lean b/Mathlib/RingTheory/DedekindDomain/SInteger.lean index b80cad635c1d7..83b4fe6cf9021 100644 --- a/Mathlib/RingTheory/DedekindDomain/SInteger.lean +++ b/Mathlib/RingTheory/DedekindDomain/SInteger.lean @@ -69,9 +69,7 @@ def integer : Subalgebra R K := theorem integer_eq : (S.integer K).toSubring = ⨅ (v) (_ : v ∉ S), (v : HeightOneSpectrum R).valuation.valuationSubring.toSubring := - SetLike.ext' <| by - -- Porting note: was `simpa only [integer, Subring.copy_eq]` - ext; simp + SetLike.ext' <| by ext; simp theorem integer_valuation_le_one (x : S.integer K) {v : HeightOneSpectrum R} (hv : v ∉ S) : v.valuation (x : K) ≤ 1 := diff --git a/Mathlib/SetTheory/Cardinal/Basic.lean b/Mathlib/SetTheory/Cardinal/Basic.lean index cf4dfeb391b9e..8c37e772ce499 100644 --- a/Mathlib/SetTheory/Cardinal/Basic.lean +++ b/Mathlib/SetTheory/Cardinal/Basic.lean @@ -1380,7 +1380,7 @@ theorem mk_finset_of_fintype [Fintype α] : #(Finset α) = 2 ^ Fintype.card α : theorem card_le_of_finset {α} (s : Finset α) : (s.card : Cardinal) ≤ #α := @mk_coe_finset _ s ▸ mk_set_le _ --- Porting note: was `simp`. LHS is not normal form. +-- Porting note (#11119): was `simp`. LHS is not normal form. -- @[simp, norm_cast] @[norm_cast] theorem natCast_pow {m n : ℕ} : (↑(m ^ n) : Cardinal) = (↑m : Cardinal) ^ (↑n : Cardinal) := by diff --git a/Mathlib/Topology/Algebra/Module/Basic.lean b/Mathlib/Topology/Algebra/Module/Basic.lean index 3a3b76145aaeb..451dff66e30df 100644 --- a/Mathlib/Topology/Algebra/Module/Basic.lean +++ b/Mathlib/Topology/Algebra/Module/Basic.lean @@ -1186,11 +1186,7 @@ def iInfKerProjEquiv {I J : Set ι} [DecidablePred fun i => i ∈ I] (hd : Disjo continuous_invFun := Continuous.subtype_mk (continuous_pi fun i => by - -- Porting note: Was `dsimp`. - change - Continuous (⇑(if h : i ∈ I then LinearMap.proj (R := R) (ι := ↥I) - (φ := fun i : ↥I => φ i) ⟨i, h⟩ else - (0 : ((i : I) → φ i) →ₗ[R] φ i))) + dsimp split_ifs <;> [apply continuous_apply; exact continuous_zero]) _ diff --git a/Mathlib/Topology/Instances/AddCircle.lean b/Mathlib/Topology/Instances/AddCircle.lean index 3295ee564f992..969b2df626d4d 100644 --- a/Mathlib/Topology/Instances/AddCircle.lean +++ b/Mathlib/Topology/Instances/AddCircle.lean @@ -591,9 +591,7 @@ homeomorphism of topological spaces. -/ def homeoIccQuot [TopologicalSpace 𝕜] [OrderTopology 𝕜] : 𝕋 ≃ₜ Quot (EndpointIdent p a) where toEquiv := equivIccQuot p a continuous_toFun := by - -- Porting note: was `simp_rw` - rw [quotientMap_quotient_mk'.continuous_iff] - simp_rw [continuous_iff_continuousAt, + simp_rw [quotientMap_quotient_mk'.continuous_iff, continuous_iff_continuousAt, continuousAt_iff_continuous_left_right] intro x; constructor on_goal 1 => erw [equivIccQuot_comp_mk_eq_toIocMod] From 402539e198cca9e794758c2ea3586a9e875d315b Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Thu, 17 Oct 2024 09:17:43 +0000 Subject: [PATCH 065/505] feat(Topology): set of generic points of generic components (#15407) Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- Mathlib/AlgebraicGeometry/FunctionField.lean | 6 +- Mathlib/Topology/Irreducible.lean | 5 + Mathlib/Topology/Sober.lean | 99 ++++++++++++++++++-- 3 files changed, 98 insertions(+), 12 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/FunctionField.lean b/Mathlib/AlgebraicGeometry/FunctionField.lean index 1b998330cf366..242c360e3de84 100644 --- a/Mathlib/AlgebraicGeometry/FunctionField.lean +++ b/Mathlib/AlgebraicGeometry/FunctionField.lean @@ -51,7 +51,7 @@ noncomputable instance [IsIntegral X] : Field X.functionField := by intro ha replace ha := ne_of_apply_ne _ ha have hs : genericPoint X ∈ RingedSpace.basicOpen _ s := by - rw [← SetLike.mem_coe, (genericPoint_spec X).mem_open_set_iff, Set.top_eq_univ, + rw [← SetLike.mem_coe, (genericPoint_spec X).mem_open_set_iff, Set.univ_inter, Set.nonempty_iff_ne_empty, Ne, ← Opens.coe_bot, ← SetLike.ext'_iff] · erw [basicOpen_eq_bot_iff] exact ha @@ -79,7 +79,7 @@ theorem genericPoint_eq_of_isOpenImmersion {X Y : Scheme} (f : X ⟶ Y) [H : IsO apply ((genericPoint_spec Y).eq _).symm convert (genericPoint_spec X).image (show Continuous f.base by fun_prop) symm - rw [eq_top_iff, Set.top_eq_univ, Set.top_eq_univ] + rw [← Set.univ_subset_iff] convert subset_closure_inter_of_isPreirreducible_of_isOpen _ H.base_open.isOpen_range _ · rw [Set.univ_inter, Set.image_univ] · apply PreirreducibleSpace.isPreirreducible_univ (X := Y) @@ -107,7 +107,7 @@ theorem genericPoint_eq_bot_of_affine (R : CommRingCat) [IsDomain R] : apply (genericPoint_spec (Spec R)).eq rw [isGenericPoint_def] rw [← PrimeSpectrum.zeroLocus_vanishingIdeal_eq_closure, PrimeSpectrum.vanishingIdeal_singleton] - rw [Set.top_eq_univ, ← PrimeSpectrum.zeroLocus_singleton_zero] + rw [← PrimeSpectrum.zeroLocus_singleton_zero] rfl instance functionField_isFractionRing_of_affine (R : CommRingCat.{u}) [IsDomain R] : diff --git a/Mathlib/Topology/Irreducible.lean b/Mathlib/Topology/Irreducible.lean index 8a868fe36faa7..17b268468df10 100644 --- a/Mathlib/Topology/Irreducible.lean +++ b/Mathlib/Topology/Irreducible.lean @@ -212,6 +212,11 @@ instance (priority := 100) {X} [Infinite X] : IrreducibleSpace (CofiniteTopology simpa only [compl_union, compl_compl] using ((hu hu').union (hv hv')).infinite_compl.nonempty toNonempty := (inferInstance : Nonempty X) +theorem irreducibleComponents_eq_singleton [IrreducibleSpace X] : + irreducibleComponents X = {univ} := + Set.ext fun _ ↦ IsGreatest.maximal_iff (s := IsIrreducible (X := X)) + ⟨IrreducibleSpace.isIrreducible_univ X, fun _ _ ↦ Set.subset_univ _⟩ + /-- A set `s` is irreducible if and only if for every finite collection of open sets all of whose members intersect `s`, `s` also intersects the intersection of the entire collection diff --git a/Mathlib/Topology/Sober.lean b/Mathlib/Topology/Sober.lean index a4717443959a5..82c12d98222d6 100644 --- a/Mathlib/Topology/Sober.lean +++ b/Mathlib/Topology/Sober.lean @@ -19,6 +19,7 @@ stated via `[QuasiSober α] [T0Space α]`. * `IsGenericPoint` : `x` is the generic point of `S` if `S` is the closure of `x`. * `QuasiSober` : A space is quasi-sober if every irreducible closed subset has a generic point. +* `genericPoints` : The set of generic points of irreducible components. -/ @@ -108,14 +109,28 @@ noncomputable def IsIrreducible.genericPoint [QuasiSober α] {S : Set α} (hS : α := (QuasiSober.sober hS.closure isClosed_closure).choose -theorem IsIrreducible.genericPoint_spec [QuasiSober α] {S : Set α} (hS : IsIrreducible S) : +theorem IsIrreducible.isGenericPoint_genericPoint_closure + [QuasiSober α] {S : Set α} (hS : IsIrreducible S) : IsGenericPoint hS.genericPoint (closure S) := (QuasiSober.sober hS.closure isClosed_closure).choose_spec +theorem IsIrreducible.isGenericPoint_genericPoint [QuasiSober α] {S : Set α} + (hS : IsIrreducible S) (hS' : IsClosed S) : + IsGenericPoint hS.genericPoint S := by + convert hS.isGenericPoint_genericPoint_closure; exact hS'.closure_eq.symm + @[simp] theorem IsIrreducible.genericPoint_closure_eq [QuasiSober α] {S : Set α} (hS : IsIrreducible S) : closure ({hS.genericPoint} : Set α) = closure S := - hS.genericPoint_spec + hS.isGenericPoint_genericPoint_closure + +theorem IsIrreducible.closure_genericPoint [QuasiSober α] {S : Set α} + (hS : IsIrreducible S) (hS' : IsClosed S) : + closure ({hS.genericPoint} : Set α) = S := + hS.isGenericPoint_genericPoint_closure.trans hS'.closure_eq + +@[deprecated (since := "2024-10-03")] +alias IsIrreducible.genericPoint_spec := IsIrreducible.isGenericPoint_genericPoint_closure variable (α) @@ -124,18 +139,18 @@ noncomputable def genericPoint [QuasiSober α] [IrreducibleSpace α] : α := (IrreducibleSpace.isIrreducible_univ α).genericPoint theorem genericPoint_spec [QuasiSober α] [IrreducibleSpace α] : - IsGenericPoint (genericPoint α) ⊤ := by - simpa using (IrreducibleSpace.isIrreducible_univ α).genericPoint_spec + IsGenericPoint (genericPoint α) univ := by + simpa using (IrreducibleSpace.isIrreducible_univ α).isGenericPoint_genericPoint_closure @[simp] theorem genericPoint_closure [QuasiSober α] [IrreducibleSpace α] : - closure ({genericPoint α} : Set α) = ⊤ := + closure ({genericPoint α} : Set α) = univ := genericPoint_spec α variable {α} theorem genericPoint_specializes [QuasiSober α] [IrreducibleSpace α] (x : α) : genericPoint α ⤳ x := - (IsIrreducible.genericPoint_spec _).specializes (by simp) + (IsIrreducible.isGenericPoint_genericPoint_closure _).specializes (by simp) attribute [local instance] specializationOrder @@ -149,7 +164,7 @@ noncomputable def irreducibleSetEquivPoints [QuasiSober α] [T0Space α] : simp only [IsIrreducible.genericPoint_closure_eq, TopologicalSpace.IrreducibleCloseds.coe_mk, closure_eq_iff_isClosed.mpr s.3] rfl - right_inv x := isIrreducible_singleton.closure.genericPoint_spec.eq + right_inv x := isIrreducible_singleton.closure.isGenericPoint_genericPoint_closure.eq (by rw [closure_closure]; exact isGenericPoint_closure) map_rel_iff' := by rintro ⟨s, hs, hs'⟩ ⟨t, ht, ht'⟩ @@ -204,12 +219,13 @@ theorem quasiSober_of_open_cover (S : Set (Set α)) (hS : ∀ s : S, IsOpen (s : h.2.preimage (hS ⟨U, hU⟩).openEmbedding_subtype_val replace H : IsIrreducible ((↑) ⁻¹' t : Set U) := ⟨⟨⟨x, hU'⟩, by simpa using hx⟩, H⟩ use H.genericPoint - have := continuous_subtype_val.closure_preimage_subset _ H.genericPoint_spec.mem + have := continuous_subtype_val.closure_preimage_subset _ H.isGenericPoint_genericPoint_closure.mem rw [h'.closure_eq] at this apply le_antisymm · apply h'.closure_subset_iff.mpr simpa using this - rw [← image_singleton, ← closure_image_closure continuous_subtype_val, H.genericPoint_spec.def] + rw [← image_singleton, ← closure_image_closure continuous_subtype_val, + H.isGenericPoint_genericPoint_closure.def] refine (subset_closure_inter_of_isPreirreducible_of_isOpen h.2 (hS ⟨U, hU⟩) ⟨x, hx, hU'⟩).trans (closure_mono ?_) rw [inter_comm t, ← Subtype.image_preimage_coe] @@ -222,3 +238,68 @@ instance (priority := 100) T2Space.quasiSober [T2Space α] : QuasiSober α where exact ⟨x, closure_singleton⟩ end Sober + +section genericPoints + +variable (α) in +/-- The set of generic points of irreducible components. -/ +def genericPoints : Set α := { x | closure {x} ∈ irreducibleComponents α } + +namespace genericPoints + +/-- The irreducible component of a generic point -/ +def component (x : genericPoints α) : irreducibleComponents α := + ⟨closure {x.1}, x.2⟩ + +lemma isGenericPoint (x : genericPoints α) : IsGenericPoint x.1 (component x).1 := rfl + +lemma component_injective [T0Space α] : Function.Injective (component (α := α)) := + fun x y e ↦ Subtype.ext ((isGenericPoint x).eq (e ▸ isGenericPoint y)) + +/-- The generic point of an irreducible component. -/ +noncomputable +def ofComponent [QuasiSober α] (x : irreducibleComponents α) : genericPoints α := + ⟨x.2.1.genericPoint, show _ ∈ irreducibleComponents α from + (x.2.1.isGenericPoint_genericPoint (isClosed_of_mem_irreducibleComponents x.1 x.2)).symm ▸ x.2⟩ + +lemma isGenericPoint_ofComponent [QuasiSober α] (x : irreducibleComponents α) : + IsGenericPoint (ofComponent x).1 x := + x.2.1.isGenericPoint_genericPoint (isClosed_of_mem_irreducibleComponents x.1 x.2) + +@[simp] +lemma component_ofComponent [QuasiSober α] (x : irreducibleComponents α) : + component (ofComponent x) = x := + Subtype.ext (isGenericPoint_ofComponent x) + +@[simp] +lemma ofComponent_component [T0Space α] [QuasiSober α] (x : genericPoints α) : + ofComponent (component x) = x := + component_injective (component_ofComponent _) + +lemma component_surjective [QuasiSober α] : Function.Surjective (component (α := α)) := + Function.HasRightInverse.surjective ⟨ofComponent, component_ofComponent⟩ + +lemma finite [T0Space α] (h : (irreducibleComponents α).Finite) : (genericPoints α).Finite := + @Finite.of_injective _ _ h _ component_injective + +/-- In a sober space, the generic points corresponds bijectively to irreducible components -/ +@[simps] +noncomputable +def equiv [T0Space α] [QuasiSober α] : genericPoints α ≃ irreducibleComponents α := + ⟨component, ofComponent, ofComponent_component, component_ofComponent⟩ + +lemma closure [QuasiSober α] : closure (genericPoints α) = Set.univ := by + refine Set.eq_univ_iff_forall.mpr fun x ↦ Set.subset_def.mp ?_ x mem_irreducibleComponent + refine (isGenericPoint_ofComponent + ⟨_, irreducibleComponent_mem_irreducibleComponents x⟩).symm.trans_subset (closure_mono ?_) + exact Set.singleton_subset_iff.mpr (ofComponent _).2 + +end genericPoints + +lemma genericPoints_eq_singleton [QuasiSober α] [IrreducibleSpace α] [T0Space α] : + genericPoints α = {genericPoint α} := by + ext x + rw [genericPoints, irreducibleComponents_eq_singleton] + exact ⟨((genericPoint_spec α).eq · |>.symm), (· ▸ genericPoint_spec α)⟩ + +end genericPoints From d8ceed52a92e5b0ba56911202c89d75114ce9d1b Mon Sep 17 00:00:00 2001 From: Christian Merten Date: Thu, 17 Oct 2024 09:39:27 +0000 Subject: [PATCH 066/505] feat(AlgebraicGeometry): smooth morphisms of schemes (#14161) We say a morphism of schemes is smooth (of relative dimension `n`) if for every point of the source there exist affine open neighborhoods of the point and its image such that the induced ring map is standard smooth (of relative dimension `n`). These notions are local on the source and the target and satisfy the standard stability properties. This contribution was created as part of the AIM workshop "Formalizing algebraic geometry" in June 2024. --- Mathlib.lean | 1 + .../Morphisms/RingHomProperties.lean | 70 ++++++++ .../AlgebraicGeometry/Morphisms/Smooth.lean | 167 ++++++++++++++++++ 3 files changed, 238 insertions(+) create mode 100644 Mathlib/AlgebraicGeometry/Morphisms/Smooth.lean diff --git a/Mathlib.lean b/Mathlib.lean index cb62c7d005329..9005c78e9dcb7 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -870,6 +870,7 @@ import Mathlib.AlgebraicGeometry.Morphisms.QuasiCompact import Mathlib.AlgebraicGeometry.Morphisms.QuasiSeparated import Mathlib.AlgebraicGeometry.Morphisms.RingHomProperties import Mathlib.AlgebraicGeometry.Morphisms.Separated +import Mathlib.AlgebraicGeometry.Morphisms.Smooth import Mathlib.AlgebraicGeometry.Morphisms.UnderlyingMap import Mathlib.AlgebraicGeometry.Morphisms.UniversallyClosed import Mathlib.AlgebraicGeometry.Noetherian diff --git a/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean b/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean index 8e9613ac56f3a..06bd0001fe507 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean @@ -163,6 +163,63 @@ theorem sourceAffineLocally_isLocal (h₁ : RingHom.RespectsIso P) rwa [IsAffineOpen.appLE_eq_away_map f (isAffineOpen_top Y) U.2, ← h₁.is_localization_away_iff] at this +open RingHom + +variable {P} {X Y : Scheme.{u}} {f : X ⟶ Y} + +/-- If `P` holds for `f` over affine opens `U₂` of `Y` and `V₂` of `X` and `U₁` (resp. `V₁`) are +open affine neighborhoods of `x` (resp. `f.base x`), then `P` also holds for `f` +over some basic open of `U₁` (resp. `V₁`). -/ +lemma exists_basicOpen_le_appLE_of_appLE_of_isAffine + (hPa : StableUnderCompositionWithLocalizationAway P) (hPl : LocalizationPreserves P) + (x : X) (U₁ : Y.affineOpens) (U₂ : Y.affineOpens) (V₁ : X.affineOpens) (V₂ : X.affineOpens) + (hx₁ : x ∈ V₁.1) (hx₂ : x ∈ V₂.1) (e₂ : V₂.1 ≤ f ⁻¹ᵁ U₂.1) (h₂ : P (f.appLE U₂ V₂ e₂)) + (hfx₁ : f.base x ∈ U₁.1) : + ∃ (r : Γ(Y, U₁)) (s : Γ(X, V₁)) (_ : x ∈ X.basicOpen s) + (e : X.basicOpen s ≤ f ⁻¹ᵁ Y.basicOpen r), P (f.appLE (Y.basicOpen r) (X.basicOpen s) e) := by + obtain ⟨r, r', hBrr', hBfx⟩ := exists_basicOpen_le_affine_inter U₁.2 U₂.2 (f.base x) + ⟨hfx₁, e₂ hx₂⟩ + have ha : IsAffineOpen (X.basicOpen (f.appLE U₂ V₂ e₂ r')) := V₂.2.basicOpen _ + have hxa : x ∈ X.basicOpen (f.appLE U₂ V₂ e₂ r') := by + simpa [Scheme.Hom.appLE, ← Scheme.preimage_basicOpen] using And.intro hx₂ (hBrr' ▸ hBfx) + obtain ⟨s, s', hBss', hBx⟩ := exists_basicOpen_le_affine_inter V₁.2 ha x ⟨hx₁, hxa⟩ + haveI := V₂.2.isLocalization_basicOpen (f.appLE U₂ V₂ e₂ r') + haveI := U₂.2.isLocalization_basicOpen r' + haveI := ha.isLocalization_basicOpen s' + have ers : X.basicOpen s ≤ f ⁻¹ᵁ Y.basicOpen r := by + rw [hBss', hBrr'] + apply le_trans (X.basicOpen_le _) + simp [Scheme.Hom.appLE] + have heq : f.appLE (Y.basicOpen r') (X.basicOpen s') (hBrr' ▸ hBss' ▸ ers) = + f.appLE (Y.basicOpen r') (X.basicOpen (f.appLE U₂ V₂ e₂ r')) (by simp [Scheme.Hom.appLE]) ≫ + algebraMap _ _ := by + simp only [Scheme.Hom.appLE, homOfLE_leOfHom, CommRingCat.comp_apply, Category.assoc] + congr + apply X.presheaf.map_comp + refine ⟨r, s, hBx, ers, ?_⟩ + · rw [f.appLE_congr _ hBrr' hBss' P, heq] + apply hPa.left _ s' _ + rw [U₂.2.appLE_eq_away_map f V₂.2] + exact hPl.away _ h₂ + +/-- If `P` holds for `f` over affine opens `U₂` of `Y` and `V₂` of `X` and `U₁` (resp. `V₁`) are +open neighborhoods of `x` (resp. `f.base x`), then `P` also holds for `f` over some affine open +`U'` of `Y` (resp. `V'` of `X`) that is contained in `U₁` (resp. `V₁`). -/ +lemma exists_affineOpens_le_appLE_of_appLE + (hPa : StableUnderCompositionWithLocalizationAway P) (hPl : LocalizationPreserves P) + (x : X) (U₁ : Y.Opens) (U₂ : Y.affineOpens) (V₁ : X.Opens) (V₂ : X.affineOpens) + (hx₁ : x ∈ V₁) (hx₂ : x ∈ V₂.1) (e₂ : V₂.1 ≤ f ⁻¹ᵁ U₂.1) (h₂ : P (f.appLE U₂ V₂ e₂)) + (hfx₁ : f.base x ∈ U₁.1) : + ∃ (U' : Y.affineOpens) (V' : X.affineOpens) (_ : U'.1 ≤ U₁) (_ : V'.1 ≤ V₁) (_ : x ∈ V'.1) + (e : V'.1 ≤ f⁻¹ᵁ U'.1), P (f.appLE U' V' e) := by + obtain ⟨r, hBr, hBfx⟩ := U₂.2.exists_basicOpen_le ⟨f.base x, hfx₁⟩ (e₂ hx₂) + obtain ⟨s, hBs, hBx⟩ := V₂.2.exists_basicOpen_le ⟨x, hx₁⟩ hx₂ + obtain ⟨r', s', hBx', e', hf'⟩ := exists_basicOpen_le_appLE_of_appLE_of_isAffine hPa hPl x + ⟨Y.basicOpen r, U₂.2.basicOpen _⟩ U₂ ⟨X.basicOpen s, V₂.2.basicOpen _⟩ V₂ hBx hx₂ e₂ h₂ hBfx + exact ⟨⟨Y.basicOpen r', (U₂.2.basicOpen _).basicOpen _⟩, + ⟨X.basicOpen s', (V₂.2.basicOpen _).basicOpen _⟩, le_trans (Y.basicOpen_le _) hBr, + le_trans (X.basicOpen_le _) hBs, hBx', e', hf'⟩ + end affineLocally /-- @@ -475,6 +532,19 @@ lemma iff_exists_appLE : P f ↔ haveI : HasRingHomProperty P Q := inst apply (isLocal_ringHomProperty P (Q := Q)).respectsIso +omit [HasRingHomProperty P Q] in +lemma locally_of_iff (hQl : LocalizationPreserves Q) + (hQa : StableUnderCompositionWithLocalizationAway Q) + (h : ∀ {X Y : Scheme.{u}} (f : X ⟶ Y), P f ↔ + ∀ (x : X), ∃ (U : Y.affineOpens) (V : X.affineOpens) (_ : x ∈ V.1) (e : V.1 ≤ f ⁻¹ᵁ U.1), + Q (f.appLE U V e)) : HasRingHomProperty P (Locally Q) where + isLocal_ringHomProperty := locally_propertyIsLocal hQl hQa + eq_affineLocally' := by + haveI : HasRingHomProperty (affineLocally (Locally Q)) (Locally Q) := + ⟨locally_propertyIsLocal hQl hQa, rfl⟩ + ext X Y f + rw [h, iff_exists_appLE_locally (P := affineLocally (Locally Q)) hQa.respectsIso] + end HasRingHomProperty end AlgebraicGeometry diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Smooth.lean b/Mathlib/AlgebraicGeometry/Morphisms/Smooth.lean new file mode 100644 index 0000000000000..31d2f234dbb6d --- /dev/null +++ b/Mathlib/AlgebraicGeometry/Morphisms/Smooth.lean @@ -0,0 +1,167 @@ +/- +Copyright (c) 2024 Christian Merten. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Christian Merten +-/ +import Mathlib.AlgebraicGeometry.Morphisms.RingHomProperties +import Mathlib.RingTheory.RingHom.StandardSmooth + +/-! + +# Smooth morphisms + +A morphism of schemes `f : X ⟶ Y` is smooth (of relative dimension `n`) if for each `x : X` there +exists an affine open neighborhood `V` of `x` and an affine open neighborhood `U` of +`f.base x` with `V ≤ f ⁻¹ᵁ U` such that the induced map `Γ(Y, U) ⟶ Γ(X, V)` is +standard smooth (of relative dimension `n`). + +In other words, smooth (resp. smooth of relative dimension `n`) for scheme morphisms is associated +to the property of ring homomorphisms `Locally IsStandardSmooth` +(resp. `Locally (IsStandardSmoothOfRelativeDimension n)`). + +## Implementation details + +- Our definition is equivalent to defining `IsSmooth` as the associated scheme morphism property of +the property of ring maps induced by `Algebra.Smooth`. The equivalence will follow from the +equivalence of `Locally IsStandardSmooth` and `Algebra.IsSmooth`, but the latter is a (hard) TODO. + +The reason why we choose the definition via `IsStandardSmooth`, is because verifying that +`Algebra.IsSmooth` is local in the sense of `RingHom.PropertyIsLocal` is a (hard) TODO. + +- The definition `RingHom.IsStandardSmooth` depends on universe levels for the generators and +relations. For morphisms of schemes we set both to `0` to avoid unnecessary complications. + +## Notes + +This contribution was created as part of the AIM workshop "Formalizing algebraic geometry" in +June 2024. + +-/ + + +noncomputable section + +open CategoryTheory + +universe t w v u + +namespace AlgebraicGeometry + +open RingHom + +variable (n m : ℕ) {X Y : Scheme.{u}} (f : X ⟶ Y) + +/-- +A morphism of schemes `f : X ⟶ Y` is smooth if for each `x : X` there +exists an affine open neighborhood `V` of `x` and an affine open neighborhood `U` of +`f.base x` with `V ≤ f ⁻¹ᵁ U` such that the induced map `Γ(Y, U) ⟶ Γ(X, V)` is +standard smooth. +-/ +@[mk_iff] +class IsSmooth : Prop where + exists_isStandardSmooth : ∀ (x : X), ∃ (U : Y.affineOpens) (V : X.affineOpens) (_ : x ∈ V.1) + (e : V.1 ≤ f ⁻¹ᵁ U.1), IsStandardSmooth.{0, 0} (f.appLE U V e) + +/-- The property of scheme morphisms `IsSmooth` is associated with the ring +homomorphism property `Locally IsStandardSmooth.{0, 0}`. -/ +instance : HasRingHomProperty @IsSmooth (Locally IsStandardSmooth.{0, 0}) := by + apply HasRingHomProperty.locally_of_iff + · exact isStandardSmooth_localizationPreserves + · exact isStandardSmooth_stableUnderCompositionWithLocalizationAway + · intro X Y f + rw [isSmooth_iff] + +/-- Being smooth is stable under composition. -/ +instance : MorphismProperty.IsStableUnderComposition @IsSmooth := + HasRingHomProperty.stableUnderComposition <| locally_stableUnderComposition + isStandardSmooth_respectsIso isStandardSmooth_localizationPreserves + isStandardSmooth_stableUnderComposition + +/-- The composition of smooth morphisms is smooth. -/ +instance isSmooth_comp {Z : Scheme.{u}} (g : Y ⟶ Z) [IsSmooth f] [IsSmooth g] : + IsSmooth (f ≫ g) := + MorphismProperty.comp_mem _ f g ‹IsSmooth f› ‹IsSmooth g› + +/-- Smooth of relative dimension `n` is stable under base change. -/ +lemma isSmooth_stableUnderBaseChange : MorphismProperty.StableUnderBaseChange @IsSmooth := + HasRingHomProperty.stableUnderBaseChange <| locally_stableUnderBaseChange + isStandardSmooth_respectsIso isStandardSmooth_stableUnderBaseChange + +/-- +A morphism of schemes `f : X ⟶ Y` is smooth of relative dimension `n` if for each `x : X` there +exists an affine open neighborhood `V` of `x` and an affine open neighborhood `U` of +`f.base x` with `V ≤ f ⁻¹ᵁ U` such that the induced map `Γ(Y, U) ⟶ Γ(X, V)` is +standard smooth of relative dimension `n`. +-/ +@[mk_iff] +class IsSmoothOfRelativeDimension : Prop where + exists_isStandardSmoothOfRelativeDimension : ∀ (x : X), ∃ (U : Y.affineOpens) + (V : X.affineOpens) (_ : x ∈ V.1) (e : V.1 ≤ f ⁻¹ᵁ U.1), + IsStandardSmoothOfRelativeDimension.{0, 0} n (f.appLE U V e) + +/-- If `f` is smooth of any relative dimension, it is smooth. -/ +lemma IsSmoothOfRelativeDimension.isSmooth [IsSmoothOfRelativeDimension n f] : IsSmooth f where + exists_isStandardSmooth x := by + obtain ⟨U, V, hx, e, hf⟩ := exists_isStandardSmoothOfRelativeDimension (n := n) (f := f) x + exact ⟨U, V, hx, e, hf.isStandardSmooth⟩ + +/-- The property of scheme morphisms `IsSmoothOfRelativeDimension n` is associated with the ring +homomorphism property `Locally (IsStandardSmoothOfRelativeDimension.{0, 0} n)`. -/ +instance : HasRingHomProperty (@IsSmoothOfRelativeDimension n) + (Locally (IsStandardSmoothOfRelativeDimension.{0, 0} n)) := by + apply HasRingHomProperty.locally_of_iff + · exact isStandardSmoothOfRelativeDimension_localizationPreserves n + · exact isStandardSmoothOfRelativeDimension_stableUnderCompositionWithLocalizationAway n + · intro X Y f + rw [isSmoothOfRelativeDimension_iff] + +/-- Smooth of relative dimension `n` is stable under base change. -/ +lemma isSmoothOfRelativeDimension_stableUnderBaseChange : + MorphismProperty.StableUnderBaseChange (@IsSmoothOfRelativeDimension n) := + HasRingHomProperty.stableUnderBaseChange <| locally_stableUnderBaseChange + isStandardSmoothOfRelativeDimension_respectsIso + (isStandardSmoothOfRelativeDimension_stableUnderBaseChange n) + +/-- Open immersions are smooth of relative dimension `0`. -/ +instance (priority := 900) [IsOpenImmersion f] : IsSmoothOfRelativeDimension 0 f := + HasRingHomProperty.of_isOpenImmersion + (locally_holdsForLocalizationAway <| + isStandardSmoothOfRelativeDimension_holdsForLocalizationAway).containsIdentities + +/-- Open immersions are smooth. -/ +instance (priority := 900) [IsOpenImmersion f] : IsSmooth f := + IsSmoothOfRelativeDimension.isSmooth 0 f + +/-- If `f` is smooth of relative dimension `n` and `g` is smooth of relative dimension +`m`, then `f ≫ g` is smooth of relative dimension `n + m`. -/ +instance isSmoothOfRelativeDimension_comp {Z : Scheme.{u}} (g : Y ⟶ Z) + [hf : IsSmoothOfRelativeDimension n f] [hg : IsSmoothOfRelativeDimension m g] : + IsSmoothOfRelativeDimension (n + m) (f ≫ g) where + exists_isStandardSmoothOfRelativeDimension x := by + obtain ⟨U₂, V₂, hfx₂, e₂, hf₂⟩ := hg.exists_isStandardSmoothOfRelativeDimension (f.base x) + obtain ⟨U₁', V₁', hx₁', e₁', hf₁'⟩ := hf.exists_isStandardSmoothOfRelativeDimension x + obtain ⟨r, s, hx₁, e₁, hf₁⟩ := exists_basicOpen_le_appLE_of_appLE_of_isAffine + (isStandardSmoothOfRelativeDimension_stableUnderCompositionWithLocalizationAway n) + (isStandardSmoothOfRelativeDimension_localizationPreserves n) + x V₂ U₁' V₁' V₁' hx₁' hx₁' e₁' hf₁' hfx₂ + have e : X.basicOpen s ≤ (f ≫ g) ⁻¹ᵁ U₂ := + le_trans e₁ <| f.preimage_le_preimage_of_le <| le_trans (Y.basicOpen_le r) e₂ + have heq : (f ≫ g).appLE U₂ (X.basicOpen s) e = g.appLE U₂ V₂ e₂ ≫ + algebraMap Γ(Y, V₂) Γ(Y, Y.basicOpen r) ≫ f.appLE (Y.basicOpen r) (X.basicOpen s) e₁ := by + rw [RingHom.algebraMap_toAlgebra, g.appLE_map_assoc, Scheme.appLE_comp_appLE] + refine ⟨U₂, ⟨X.basicOpen s, V₁'.2.basicOpen s⟩, hx₁, e, heq ▸ ?_⟩ + apply IsStandardSmoothOfRelativeDimension.comp ?_ hf₂ + haveI : IsLocalization.Away r Γ(Y, Y.basicOpen r) := V₂.2.isLocalization_basicOpen r + exact (isStandardSmoothOfRelativeDimension_stableUnderCompositionWithLocalizationAway n).right + _ r _ hf₁ + +instance {Z : Scheme.{u}} (g : Y ⟶ Z) [IsSmoothOfRelativeDimension 0 f] + [IsSmoothOfRelativeDimension 0 g] : + IsSmoothOfRelativeDimension 0 (f ≫ g) := + inferInstanceAs <| IsSmoothOfRelativeDimension (0 + 0) (f ≫ g) + +/-- Smooth of relative dimension `0` is stable under composition. -/ +instance : MorphismProperty.IsStableUnderComposition (@IsSmoothOfRelativeDimension 0) where + comp_mem _ _ _ _ := inferInstance + +end AlgebraicGeometry From c199a0e22d7561cee51596b0230955f96f882153 Mon Sep 17 00:00:00 2001 From: Christian Merten Date: Thu, 17 Oct 2024 09:39:28 +0000 Subject: [PATCH 067/505] feat(AlgebraicGeometry): closed immersions with affine target (#15038) Shows that closed immersions are affine morphisms inducing affine locally surjective maps on sections. --- Mathlib/Algebra/Category/Ring/Basic.lean | 10 ++ Mathlib/Algebra/GroupWithZero/Basic.lean | 5 + Mathlib/AlgebraicGeometry/AffineScheme.lean | 15 ++ .../Morphisms/ClosedImmersion.lean | 136 +++++++++++++++++- .../Morphisms/QuasiCompact.lean | 14 ++ .../Morphisms/QuasiSeparated.lean | 26 ++++ Mathlib/AlgebraicGeometry/ResidueField.lean | 15 ++ Mathlib/AlgebraicGeometry/Scheme.lean | 8 ++ Mathlib/Geometry/RingedSpace/Basic.lean | 15 ++ .../LocallyRingedSpace/ResidueField.lean | 6 + Mathlib/RingTheory/Localization/Defs.lean | 18 +++ Mathlib/Topology/Category/TopCat/Basic.lean | 12 ++ Mathlib/Topology/Homeomorph.lean | 10 ++ 13 files changed, 286 insertions(+), 4 deletions(-) diff --git a/Mathlib/Algebra/Category/Ring/Basic.lean b/Mathlib/Algebra/Category/Ring/Basic.lean index 718c2490097f5..801518f945a60 100644 --- a/Mathlib/Algebra/Category/Ring/Basic.lean +++ b/Mathlib/Algebra/Category/Ring/Basic.lean @@ -584,6 +584,16 @@ theorem commRingIsoToRingEquiv_symm_toRingHom {X Y : CommRingCat} (i : X ≅ Y) ext rfl +@[simp] +lemma commRingIsoToRingEquiv_apply {X Y : CommRingCat} (i : X ≅ Y) (x : X) : + i.commRingCatIsoToRingEquiv x = i.hom x := + rfl + +@[simp] +lemma commRingIsoToRingEquiv_symm_apply {X Y : CommRingCat} (i : X ≅ Y) (y : Y) : + i.commRingCatIsoToRingEquiv.symm y = i.inv y := + rfl + end CategoryTheory.Iso /-- Ring equivalences between `RingCat`s are the same as (isomorphic to) isomorphisms in diff --git a/Mathlib/Algebra/GroupWithZero/Basic.lean b/Mathlib/Algebra/GroupWithZero/Basic.lean index da8ba43a46d39..64fdcf4a2b337 100644 --- a/Mathlib/Algebra/GroupWithZero/Basic.lean +++ b/Mathlib/Algebra/GroupWithZero/Basic.lean @@ -156,6 +156,11 @@ lemma ne_zero_pow (hn : n ≠ 0) (ha : a ^ n ≠ 0) : a ≠ 0 := by rintro rfl; lemma zero_pow_eq_zero [Nontrivial M₀] : (0 : M₀) ^ n = 0 ↔ n ≠ 0 := ⟨by rintro h rfl; simp at h, zero_pow⟩ +lemma pow_mul_eq_zero_of_le {a b : M₀} {m n : ℕ} (hmn : m ≤ n) + (h : a ^ m * b = 0) : a ^ n * b = 0 := by + rw [show n = n - m + m by omega, pow_add, mul_assoc, h] + simp + variable [NoZeroDivisors M₀] lemma pow_eq_zero : ∀ {n}, a ^ n = 0 → a = 0 diff --git a/Mathlib/AlgebraicGeometry/AffineScheme.lean b/Mathlib/AlgebraicGeometry/AffineScheme.lean index 731a32691323c..0859e4051faef 100644 --- a/Mathlib/AlgebraicGeometry/AffineScheme.lean +++ b/Mathlib/AlgebraicGeometry/AffineScheme.lean @@ -651,6 +651,15 @@ theorem isLocalization_stalk (x : U) : subst this exact hU.isLocalization_stalk' y hx +lemma stalkMap_injective (f : X ⟶ Y) {U : Opens Y} (hU : IsAffineOpen U) (x : X) + (hx : f.base x ∈ U) + (h : ∀ g, f.stalkMap x (Y.presheaf.germ U (f.base x) hx g) = 0 → + Y.presheaf.germ U (f.base x) hx g = 0) : + Function.Injective (f.stalkMap x) := by + letI := Y.presheaf.algebra_section_stalk ⟨f.base x, hx⟩ + apply (hU.isLocalization_stalk ⟨f.base x, hx⟩).injective_of_map_algebraMap_zero + exact h + /-- The basic open set of a section `f` on an affine open as an `X.affineOpens`. -/ @[simps] def _root_.AlgebraicGeometry.Scheme.affineBasicOpen @@ -701,6 +710,12 @@ theorem self_le_basicOpen_union_iff (s : Set Γ(X, U)) : end IsAffineOpen +lemma stalkMap_injective_of_isAffine {X Y : Scheme} (f : X ⟶ Y) [IsAffine Y] (x : X) + (h : ∀ g, f.stalkMap x (Y.presheaf.Γgerm (f.base x) g) = 0 → + Y.presheaf.Γgerm (f.base x) g = 0) : + Function.Injective (f.stalkMap x) := + (isAffineOpen_top Y).stalkMap_injective f x trivial h + /-- Given a spanning set of `Γ(X, U)`, the corresponding basic open sets cover `U`. See `IsAffineOpen.basicOpen_union_eq_self_iff` for the inverse direction for affine open sets. diff --git a/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean b/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean index 592f5829e66fa..c78e3b86d608d 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean @@ -3,8 +3,11 @@ Copyright (c) 2023 Jonas van der Schaaf. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Amelia Livingston, Christian Merten, Jonas van der Schaaf -/ -import Mathlib.AlgebraicGeometry.Morphisms.QuasiCompact import Mathlib.AlgebraicGeometry.Morphisms.Preimmersion +import Mathlib.AlgebraicGeometry.Morphisms.QuasiSeparated +import Mathlib.AlgebraicGeometry.Morphisms.RingHomProperties +import Mathlib.AlgebraicGeometry.ResidueField +import Mathlib.RingTheory.RingHom.Surjective /-! @@ -19,8 +22,6 @@ is a closed immersion and the induced morphisms of stalks are all surjective. ## TODO -* Show closed immersions of affines are induced by surjective ring maps -* Show closed immersions are stable under pullback * Show closed immersions are precisely the proper monomorphisms * Define closed immersions of locally ringed spaces, where we also assume that the kernel of `O_X → f_*O_Y` is locally generated by sections as an `O_X`-module, and relate it to this file. See @@ -30,7 +31,7 @@ is a closed immersion and the induced morphisms of stalks are all surjective. universe v u -open CategoryTheory +open CategoryTheory TopologicalSpace Opposite namespace AlgebraicGeometry @@ -137,8 +138,135 @@ instance {X Y : Scheme} (f : X ⟶ Y) [IsClosedImmersion f] : QuasiCompact f whe end IsClosedImmersion +section Affine + +variable {X Y : Scheme.{u}} [IsAffine Y] {f : X ⟶ Y} + +open IsClosedImmersion LocallyRingedSpace + +/-- If `f : X ⟶ Y` is a morphism of schemes with quasi-compact source and affine target, `f` +has a closed image and `f` induces an injection on global sections, then +`f` is surjective. -/ +lemma surjective_of_isClosed_range_of_injective [CompactSpace X] + (hfcl : IsClosed (Set.range f.base)) (hfinj : Function.Injective (f.app ⊤)) : + Function.Surjective f.base := by + obtain ⟨I, hI⟩ := (Scheme.eq_zeroLocus_of_isClosed_of_isAffine Y (Set.range f.base)).mp hfcl + let 𝒰 : X.OpenCover := X.affineCover.finiteSubcover + haveI (i : 𝒰.J) : IsAffine (𝒰.obj i) := Scheme.isAffine_affineCover X _ + apply Set.range_iff_surjective.mp + apply hI ▸ (Scheme.zeroLocus_eq_top_iff_subset_nilradical _).mpr + intro s hs + simp only [AddSubsemigroup.mem_carrier, AddSubmonoid.mem_toSubsemigroup, + Submodule.mem_toAddSubmonoid, SetLike.mem_coe, mem_nilradical, ← IsNilpotent.map_iff hfinj] + refine Scheme.isNilpotent_of_isNilpotent_cover _ 𝒰 (fun i ↦ ?_) + rw [Scheme.isNilpotent_iff_basicOpen_eq_bot] + rw [Scheme.basicOpen_eq_bot_iff_forall_evaluation_eq_zero] + intro x + suffices h : f.base ((𝒰.map i).base x.val) ∉ Y.basicOpen s by + erw [← Scheme.Γevaluation_naturality_apply (𝒰.map i ≫ f)] + simpa only [Scheme.comp_base, TopCat.coe_comp, Function.comp_apply, + Scheme.residueFieldMap_comp, CommRingCat.comp_apply, map_eq_zero, + Scheme.evaluation_eq_zero_iff_not_mem_basicOpen] + exact (Y.mem_zeroLocus_iff I _).mp (hI ▸ Set.mem_range_self ((𝒰.map i).base x.val)) s hs + +/-- If `f : X ⟶ Y` is open, injective, `X` is quasi-compact and `Y` is affine, then `f` is stalkwise +injective if it is injective on global sections. -/ +lemma stalkMap_injective_of_isOpenMap_of_injective [CompactSpace X] + (hfopen : IsOpenMap f.base) (hfinj₁ : Function.Injective f.base) + (hfinj₂ : Function.Injective (f.app ⊤)) (x : X) : + Function.Injective (f.stalkMap x) := by + let φ : Γ(Y, ⊤) ⟶ Γ(X, ⊤) := f.app ⊤ + let 𝒰 : X.OpenCover := X.affineCover.finiteSubcover + have (i : 𝒰.J) : IsAffine (𝒰.obj i) := Scheme.isAffine_affineCover X _ + let res (i : 𝒰.J) : Γ(X, ⊤) ⟶ Γ(𝒰.obj i, ⊤) := (𝒰.map i).app ⊤ + refine stalkMap_injective_of_isAffine _ _ (fun (g : Γ(Y, ⊤)) h ↦ ?_) + rw [TopCat.Presheaf.Γgerm, Scheme.stalkMap_germ_apply] at h + obtain ⟨U, w, (hx : x ∈ U), hg⟩ := + X.toRingedSpace.exists_res_eq_zero_of_germ_eq_zero ⊤ (φ g) ⟨x, trivial⟩ h + obtain ⟨_, ⟨s, rfl⟩, hyv, bsle⟩ := Opens.isBasis_iff_nbhd.mp (isBasis_basicOpen Y) + (show f.base x ∈ ⟨f.base '' U.carrier, hfopen U.carrier U.is_open'⟩ from ⟨x, by simpa⟩) + let W (i : 𝒰.J) : TopologicalSpace.Opens (𝒰.obj i) := (𝒰.obj i).basicOpen ((res i) (φ s)) + have hwle (i : 𝒰.J) : W i ≤ (𝒰.map i)⁻¹ᵁ U := by + show ((𝒰.obj i).basicOpen ((𝒰.map i ≫ f).app ⊤ s)) ≤ _ + rw [← Scheme.preimage_basicOpen, Scheme.comp_coeBase, Opens.map_comp_obj] + refine Scheme.Hom.preimage_le_preimage_of_le _ + (le_trans (f.preimage_le_preimage_of_le bsle) (le_of_eq ?_)) + simp [Set.preimage_image_eq _ hfinj₁] + have h0 (i : 𝒰.J) : (𝒰.map i).appLE _ (W i) (by simp) (φ g) = 0 := by + rw [← Scheme.Hom.appLE_map _ _ (homOfLE <| hwle i).op, ← Scheme.Hom.map_appLE _ le_rfl w.op] + simp only [CommRingCat.coe_comp_of, RingHom.coe_comp, Function.comp_apply] + erw [hg] + simp only [map_zero] + have h1 (i : 𝒰.J) : ∃ n, (res i) (φ (s ^ n * g)) = 0 := by + obtain ⟨n, hn⟩ := exists_of_res_zero_of_qcqs_of_top (s := ((res i) (φ s))) (h0 i) + exact ⟨n, by rwa [map_mul, map_mul, map_pow, map_pow]⟩ + have h2 : ∃ n, ∀ i, (res i) (φ (s ^ n * g)) = 0 := by + choose fn hfn using h1 + refine ⟨Finset.sup Finset.univ fn, fun i ↦ ?_⟩ + rw [map_mul, map_pow, map_mul, map_pow] + simp only [map_mul, map_pow, map_mul, map_pow] at hfn + apply pow_mul_eq_zero_of_le (Finset.le_sup (Finset.mem_univ i)) (hfn i) + obtain ⟨n, hn⟩ := h2 + apply germ_eq_zero_of_pow_mul_eq_zero (U := ⊤) ⟨f.base x, trivial⟩ hyv + rw [RingHom.injective_iff_ker_eq_bot, RingHom.ker_eq_bot_iff_eq_zero] at hfinj₂ + exact hfinj₂ _ (Scheme.zero_of_zero_cover _ _ hn) + +namespace IsClosedImmersion + +/-- If `f` is a closed immersion with affine target such that the induced map on global +sections is injective, `f` is an isomorphism. -/ +theorem isIso_of_injective_of_isAffine [IsClosedImmersion f] + (hf : Function.Injective (f.app ⊤)) : IsIso f := (isIso_iff_stalk_iso f).mpr <| + have : CompactSpace X := (closedEmbedding f).compactSpace + have hiso : IsIso f.base := TopCat.isIso_of_bijective_of_isClosedMap _ + ⟨(closedEmbedding f).inj, + surjective_of_isClosed_range_of_injective ((closedEmbedding f).isClosed_range) hf⟩ + ((closedEmbedding f).isClosedMap) + ⟨hiso, fun x ↦ (ConcreteCategory.isIso_iff_bijective _).mpr + ⟨stalkMap_injective_of_isOpenMap_of_injective ((TopCat.homeoOfIso (asIso f.base)).isOpenMap) + (closedEmbedding f).inj hf _, f.stalkMap_surjective x⟩⟩ + +variable (f) + +/-- If `f` is a closed immersion with affine target, the source is affine and +the induced map on global sections is surjective. -/ +theorem isAffine_surjective_of_isAffine [IsClosedImmersion f] : + IsAffine X ∧ Function.Surjective (f.app ⊤) := by + haveI i : IsClosedImmersion f := inferInstance + rw [← affineTargetImageFactorization_comp f] at i ⊢ + haveI := of_surjective_of_isAffine (affineTargetImageInclusion f) + (affineTargetImageInclusion_app_surjective f) + haveI := IsClosedImmersion.of_comp (affineTargetImageFactorization f) + (affineTargetImageInclusion f) + haveI := isIso_of_injective_of_isAffine (affineTargetImageFactorization_app_injective f) + exact ⟨isAffine_of_isIso (affineTargetImageFactorization f), + (ConcreteCategory.bijective_of_isIso + ((affineTargetImageFactorization f).app ⊤)).surjective.comp <| + affineTargetImageInclusion_app_surjective f⟩ + +end IsClosedImmersion + +end Affine + /-- Being a closed immersion is local at the target. -/ instance IsClosedImmersion.isLocalAtTarget : IsLocalAtTarget @IsClosedImmersion := eq_inf ▸ inferInstance +/-- On morphisms with affine target, being a closed immersion is precisely having affine source +and being surjective on global sections. -/ +instance IsClosedImmersion.hasAffineProperty : HasAffineProperty @IsClosedImmersion + (fun X Y f ↦ IsAffine X ∧ Function.Surjective (f.app ⊤)) := by + convert HasAffineProperty.of_isLocalAtTarget @IsClosedImmersion + refine ⟨fun ⟨h₁, h₂⟩ ↦ of_surjective_of_isAffine _ h₂, by apply isAffine_surjective_of_isAffine⟩ + +/-- Being a closed immersion is stable under base change. -/ +lemma IsClosedImmersion.stableUnderBaseChange : + MorphismProperty.StableUnderBaseChange @IsClosedImmersion := by + apply HasAffineProperty.stableUnderBaseChange + haveI := HasAffineProperty.isLocal_affineProperty @IsClosedImmersion + apply AffineTargetMorphismProperty.StableUnderBaseChange.mk + intro X Y S _ _ f g ⟨ha, hsurj⟩ + exact ⟨inferInstance, RingHom.surjective_stableUnderBaseChange.pullback_fst_app_top _ + RingHom.surjective_respectsIso f _ hsurj⟩ + end AlgebraicGeometry diff --git a/Mathlib/AlgebraicGeometry/Morphisms/QuasiCompact.lean b/Mathlib/AlgebraicGeometry/Morphisms/QuasiCompact.lean index 2e50b205e0b3e..13bfe53b9c827 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/QuasiCompact.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/QuasiCompact.lean @@ -272,6 +272,13 @@ lemma Scheme.isNilpotent_iff_basicOpen_eq_bot_of_isCompact {X : Scheme.{u}} rw [mul_one] at hn use n +/-- A global section of a quasi-compact scheme is nilpotent if and only if its associated +basic open is empty. -/ +lemma Scheme.isNilpotent_iff_basicOpen_eq_bot {X : Scheme.{u}} + [CompactSpace X] (f : Γ(X, ⊤)) : + IsNilpotent f ↔ X.basicOpen f = ⊥ := + isNilpotent_iff_basicOpen_eq_bot_of_isCompact (U := ⊤) (CompactSpace.isCompact_univ) f + /-- The zero locus of a set of sections over a compact open of a scheme is `X` if and only if `s` is contained in the nilradical of `Γ(X, U)`. -/ lemma Scheme.zeroLocus_eq_top_iff_subset_nilradical_of_isCompact {X : Scheme.{u}} {U : X.Opens} @@ -280,4 +287,11 @@ lemma Scheme.zeroLocus_eq_top_iff_subset_nilradical_of_isCompact {X : Scheme.{u} simp [Scheme.zeroLocus_def, ← Scheme.isNilpotent_iff_basicOpen_eq_bot_of_isCompact hU, ← mem_nilradical, Set.subset_def] +/-- The zero locus of a set of sections over a compact open of a scheme is `X` if and only if +`s` is contained in the nilradical of `Γ(X, U)`. -/ +lemma Scheme.zeroLocus_eq_top_iff_subset_nilradical {X : Scheme.{u}} + [CompactSpace X] (s : Set Γ(X, ⊤)) : + X.zeroLocus s = ⊤ ↔ s ⊆ nilradical Γ(X, ⊤) := + zeroLocus_eq_top_iff_subset_nilradical_of_isCompact (U := ⊤) (CompactSpace.isCompact_univ) s + end AlgebraicGeometry diff --git a/Mathlib/AlgebraicGeometry/Morphisms/QuasiSeparated.lean b/Mathlib/AlgebraicGeometry/Morphisms/QuasiSeparated.lean index be7e03dfc20b1..fffb3c4c1f9e1 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/QuasiSeparated.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/QuasiSeparated.lean @@ -347,6 +347,32 @@ theorem is_localization_basicOpen_of_qcqs {X : Scheme} {U : X.Opens} (hU : IsCom refine ⟨⟨_, n, rfl⟩, ?_⟩ simpa [mul_comm z] using e +lemma exists_of_res_eq_of_qcqs {X : Scheme.{u}} {U : TopologicalSpace.Opens X} + (hU : IsCompact U.carrier) (hU' : IsQuasiSeparated U.carrier) + {f g s : Γ(X, U)} (hfg : f |_ X.basicOpen s = g |_ X.basicOpen s) : + ∃ n, s ^ n * f = s ^ n * g := by + obtain ⟨n, hc⟩ := (is_localization_basicOpen_of_qcqs hU hU' s).exists_of_eq s hfg + use n + +lemma exists_of_res_eq_of_qcqs_of_top {X : Scheme.{u}} [CompactSpace X] [QuasiSeparatedSpace X] + {f g s : Γ(X, ⊤)} (hfg : f |_ X.basicOpen s = g |_ X.basicOpen s) : + ∃ n, s ^ n * f = s ^ n * g := + exists_of_res_eq_of_qcqs (U := ⊤) CompactSpace.isCompact_univ isQuasiSeparated_univ hfg + +lemma exists_of_res_zero_of_qcqs {X : Scheme.{u}} {U : TopologicalSpace.Opens X} + (hU : IsCompact U.carrier) (hU' : IsQuasiSeparated U.carrier) + {f s : Γ(X, U)} (hf : f |_ X.basicOpen s = 0) : + ∃ n, s ^ n * f = 0 := by + suffices h : ∃ n, s ^ n * f = s ^ n * 0 by + simpa using h + apply exists_of_res_eq_of_qcqs hU hU' + simpa + +lemma exists_of_res_zero_of_qcqs_of_top {X : Scheme} [CompactSpace X] [QuasiSeparatedSpace X] + {f s : Γ(X, ⊤)} (hf : f |_ X.basicOpen s = 0) : + ∃ n, s ^ n * f = 0 := + exists_of_res_zero_of_qcqs (U := ⊤) CompactSpace.isCompact_univ isQuasiSeparated_univ hf + /-- If `U` is qcqs, then `Γ(X, D(f)) ≃ Γ(X, U)_f` for every `f : Γ(X, U)`. This is known as the **Qcqs lemma** in [R. Vakil, *The rising sea*][RisingSea]. -/ theorem isIso_ΓSpec_adjunction_unit_app_basicOpen {X : Scheme} [CompactSpace X] diff --git a/Mathlib/AlgebraicGeometry/ResidueField.lean b/Mathlib/AlgebraicGeometry/ResidueField.lean index dc139a1ac5f98..a429de3ad3174 100644 --- a/Mathlib/AlgebraicGeometry/ResidueField.lean +++ b/Mathlib/AlgebraicGeometry/ResidueField.lean @@ -77,6 +77,10 @@ lemma evaluation_ne_zero_iff_mem_basicOpen (x : X) (hx : x ∈ U) (f : Γ(X, U)) X.evaluation U x hx f ≠ 0 ↔ x ∈ X.basicOpen f := by simp +lemma basicOpen_eq_bot_iff_forall_evaluation_eq_zero (f : X.presheaf.obj (op U)) : + X.basicOpen f = ⊥ ↔ ∀ (x : U), X.evaluation U x x.property f = 0 := + X.toLocallyRingedSpace.basicOpen_eq_bot_iff_forall_evaluation_eq_zero f + variable {X Y : Scheme.{u}} (f : X ⟶ Y) @@ -121,6 +125,17 @@ lemma evaluation_naturality_apply {V : Opens Y} (x : X) (hx : f.base x ∈ V) (s X.evaluation (f ⁻¹ᵁ V) x hx (f.app V s) := LocallyRingedSpace.evaluation_naturality_apply f.1 ⟨x, hx⟩ s +@[reassoc] +lemma Γevaluation_naturality (x : X) : + Y.Γevaluation (f.base x) ≫ f.residueFieldMap x = + f.c.app (op ⊤) ≫ X.Γevaluation x := + LocallyRingedSpace.Γevaluation_naturality f.toLRSHom x + +lemma Γevaluation_naturality_apply (x : X) (a : Y.presheaf.obj (op ⊤)) : + f.residueFieldMap x (Y.Γevaluation (f.base x) a) = + X.Γevaluation x (f.c.app (op ⊤) a) := + LocallyRingedSpace.Γevaluation_naturality_apply f.toLRSHom x a + instance [IsOpenImmersion f] (x) : IsIso (f.residueFieldMap x) := (LocalRing.ResidueField.mapEquiv (asIso (f.stalkMap x)).commRingCatIsoToRingEquiv).toCommRingCatIso.isIso_hom diff --git a/Mathlib/AlgebraicGeometry/Scheme.lean b/Mathlib/AlgebraicGeometry/Scheme.lean index 546fea88d15d5..b958a149bf9ec 100644 --- a/Mathlib/AlgebraicGeometry/Scheme.lean +++ b/Mathlib/AlgebraicGeometry/Scheme.lean @@ -604,6 +604,14 @@ theorem Scheme.Spec_map_presheaf_map_eqToHom {X : Scheme} {U V : X.Opens} (h : U refine (Scheme.congr_app this _).trans ?_ simp [eqToHom_map] +lemma germ_eq_zero_of_pow_mul_eq_zero {X : Scheme.{u}} {U : Opens X} (x : U) {f s : Γ(X, U)} + (hx : x.val ∈ X.basicOpen s) {n : ℕ} (hf : s ^ n * f = 0) : X.presheaf.germ U x x.2 f = 0 := by + rw [Scheme.mem_basicOpen] at hx + have hu : IsUnit (X.presheaf.germ _ x x.2 (s ^ n)) := by + rw [map_pow] + exact IsUnit.pow n hx + rw [← hu.mul_right_eq_zero, ← map_mul, hf, map_zero] + @[reassoc (attr := simp)] lemma Scheme.iso_hom_base_inv_base {X Y : Scheme.{u}} (e : X ≅ Y) : e.hom.base ≫ e.inv.base = 𝟙 _ := diff --git a/Mathlib/Geometry/RingedSpace/Basic.lean b/Mathlib/Geometry/RingedSpace/Basic.lean index a1960cb7b7970..4a89dde97d9c9 100644 --- a/Mathlib/Geometry/RingedSpace/Basic.lean +++ b/Mathlib/Geometry/RingedSpace/Basic.lean @@ -41,12 +41,27 @@ namespace RingedSpace open SheafedSpace +@[simp] +lemma res_zero {X : RingedSpace.{u}} {U V : TopologicalSpace.Opens X} + (hUV : U ≤ V) : (0 : X.presheaf.obj (op V)) |_ U = 0 := + map_zero _ + variable (X : RingedSpace) -- Porting note (#10670): this was not necessary in mathlib3 instance : CoeSort RingedSpace Type* where coe X := X.carrier +/-- If the germ of a section `f` is zero in the stalk at `x`, then `f` is zero on some neighbourhood +around `x`. -/ +lemma exists_res_eq_zero_of_germ_eq_zero (U : Opens X) (f : X.presheaf.obj (op U)) (x : U) + (h : X.presheaf.germ U x.val x.property f = 0) : + ∃ (V : Opens X) (i : V ⟶ U) (_ : x.1 ∈ V), X.presheaf.map i.op f = 0 := by + have h1 : X.presheaf.germ U x.val x.property f = X.presheaf.germ U x.val x.property 0 := by simpa + obtain ⟨V, hv, i, _, hv4⟩ := TopCat.Presheaf.germ_eq X.presheaf x.1 x.2 x.2 f 0 h1 + use V, i, hv + simpa using hv4 + /-- If the germ of a section `f` is a unit in the stalk at `x`, then `f` must be a unit on some small neighborhood around `x`. diff --git a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/ResidueField.lean b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/ResidueField.lean index b09ad64961538..89243a71df33e 100644 --- a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/ResidueField.lean +++ b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/ResidueField.lean @@ -68,6 +68,12 @@ lemma evaluation_ne_zero_iff_mem_basicOpen (x : U) (f : X.presheaf.obj (op U)) : X.evaluation x f ≠ 0 ↔ x.val ∈ X.toRingedSpace.basicOpen f := by simp +lemma basicOpen_eq_bot_iff_forall_evaluation_eq_zero (f : X.presheaf.obj (op U)) : + X.toRingedSpace.basicOpen f = ⊥ ↔ ∀ (x : U), X.evaluation x f = 0 := by + simp only [evaluation_eq_zero_iff_not_mem_basicOpen, Subtype.forall] + exact ⟨fun h ↦ h ▸ fun a _ hc ↦ hc, + fun h ↦ eq_bot_iff.mpr <| fun a ha ↦ h a (X.toRingedSpace.basicOpen_le f ha) ha⟩ + @[simp] lemma Γevaluation_eq_zero_iff_not_mem_basicOpen (x : X) (f : X.presheaf.obj (op ⊤)) : X.Γevaluation x f = 0 ↔ x ∉ X.toRingedSpace.basicOpen f := diff --git a/Mathlib/RingTheory/Localization/Defs.lean b/Mathlib/RingTheory/Localization/Defs.lean index 4206bfa0fece7..4a4ae21fb0303 100644 --- a/Mathlib/RingTheory/Localization/Defs.lean +++ b/Mathlib/RingTheory/Localization/Defs.lean @@ -528,6 +528,14 @@ theorem lift_injective_iff : Injective (lift hg : S → P) ↔ ∀ x y, algebraMap R S x = algebraMap R S y ↔ g x = g y := (toLocalizationMap M S).lift_injective_iff hg +variable (M) in +include M in +lemma injective_iff_map_algebraMap_eq {T} [CommRing T] (f : S →+* T) : + Function.Injective f ↔ ∀ x y, + algebraMap R S x = algebraMap R S y ↔ f (algebraMap R S x) = f (algebraMap R S y) := by + rw [← IsLocalization.lift_of_comp (M := M) f, IsLocalization.lift_injective_iff] + simp + section Map variable {T : Submonoid P} {Q : Type*} [CommSemiring Q] @@ -828,6 +836,16 @@ namespace IsLocalization variable {K : Type*} [IsLocalization M S] +include M in +lemma injective_of_map_algebraMap_zero {T} [CommRing T] (f : S →+* T) + (h : ∀ x, f (algebraMap R S x) = 0 → algebraMap R S x = 0) : + Function.Injective f := by + rw [IsLocalization.injective_iff_map_algebraMap_eq M] + refine fun x y ↦ ⟨fun hz ↦ hz ▸ rfl, fun hz ↦ ?_⟩ + rw [← sub_eq_zero, ← map_sub, ← map_sub] at hz + apply h at hz + rwa [map_sub, sub_eq_zero] at hz + theorem to_map_eq_zero_iff {x : R} (hM : M ≤ nonZeroDivisors R) : algebraMap R S x = 0 ↔ x = 0 := by rw [← (algebraMap R S).map_zero] constructor <;> intro h diff --git a/Mathlib/Topology/Category/TopCat/Basic.lean b/Mathlib/Topology/Category/TopCat/Basic.lean index ab44aa713d2eb..468168d5c7fb5 100644 --- a/Mathlib/Topology/Category/TopCat/Basic.lean +++ b/Mathlib/Topology/Category/TopCat/Basic.lean @@ -161,6 +161,18 @@ theorem of_homeoOfIso {X Y : TopCat.{u}} (f : X ≅ Y) : isoOfHomeo (homeoOfIso ext rfl +lemma isIso_of_bijective_of_isOpenMap {X Y : TopCat.{u}} (f : X ⟶ Y) + (hfbij : Function.Bijective f) (hfcl : IsOpenMap f) : IsIso f := + let e : X ≃ₜ Y := Homeomorph.homeomorphOfContinuousOpen + (Equiv.ofBijective f hfbij) f.continuous hfcl + inferInstanceAs <| IsIso (TopCat.isoOfHomeo e).hom + +lemma isIso_of_bijective_of_isClosedMap {X Y : TopCat.{u}} (f : X ⟶ Y) + (hfbij : Function.Bijective f) (hfcl : IsClosedMap f) : IsIso f := + let e : X ≃ₜ Y := Homeomorph.homeomorphOfContinuousClosed + (Equiv.ofBijective f hfbij) f.continuous hfcl + inferInstanceAs <| IsIso (TopCat.isoOfHomeo e).hom + -- Porting note: simpNF requested partially simped version below theorem openEmbedding_iff_comp_isIso {X Y Z : TopCat} (f : X ⟶ Y) (g : Y ⟶ Z) [IsIso g] : OpenEmbedding (f ≫ g) ↔ OpenEmbedding f := diff --git a/Mathlib/Topology/Homeomorph.lean b/Mathlib/Topology/Homeomorph.lean index 9daf8a9d2be4c..09be530f7a37a 100644 --- a/Mathlib/Topology/Homeomorph.lean +++ b/Mathlib/Topology/Homeomorph.lean @@ -417,6 +417,16 @@ def homeomorphOfContinuousOpen (e : X ≃ Y) (h₁ : Continuous e) (h₂ : IsOpe apply e.image_eq_preimage toEquiv := e +/-- If a bijective map `e : X ≃ Y` is continuous and closed, then it is a homeomorphism. -/ +def homeomorphOfContinuousClosed (e : X ≃ Y) (h₁ : Continuous e) (h₂ : IsClosedMap e) : X ≃ₜ Y where + continuous_toFun := h₁ + continuous_invFun := by + rw [continuous_iff_isClosed] + intro s hs + convert ← h₂ s hs using 1 + apply e.image_eq_preimage + toEquiv := e + @[simp] theorem homeomorphOfContinuousOpen_apply (e : X ≃ Y) (h₁ : Continuous e) (h₂ : IsOpenMap e) : ⇑(homeomorphOfContinuousOpen e h₁ h₂) = e := rfl From 035f251efa085d58f145e0d72477526d100c271e Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Thu, 17 Oct 2024 09:39:30 +0000 Subject: [PATCH 068/505] chore: clean up "was rw" porting notes (#17822) If they can be fixed, do so. If they are about changing `rw` to `erw`, refer to #11224 instead. Otherwise, only keep the note if the new proof is actually less idiomatic. None of the original "was rw" notes remain, so: Closes: #10691 --- Mathlib/Algebra/MonoidAlgebra/Defs.lean | 2 -- Mathlib/Algebra/Order/ToIntervalMod.lean | 2 +- Mathlib/Algebra/Polynomial/Laurent.lean | 2 +- Mathlib/AlgebraicGeometry/OpenImmersion.lean | 4 ++-- Mathlib/Data/DFinsupp/Interval.lean | 1 - Mathlib/LinearAlgebra/Alternating/DomCoprod.lean | 2 +- Mathlib/LinearAlgebra/CliffordAlgebra/EvenEquiv.lean | 2 +- Mathlib/Probability/Kernel/MeasurableIntegral.lean | 2 +- Mathlib/RingTheory/Localization/Ideal.lean | 4 +--- 9 files changed, 8 insertions(+), 13 deletions(-) diff --git a/Mathlib/Algebra/MonoidAlgebra/Defs.lean b/Mathlib/Algebra/MonoidAlgebra/Defs.lean index 5f5174af5e945..7125b7d636d97 100644 --- a/Mathlib/Algebra/MonoidAlgebra/Defs.lean +++ b/Mathlib/Algebra/MonoidAlgebra/Defs.lean @@ -633,7 +633,6 @@ def singleOneRingHom [Semiring k] [MulOneClass G] : k →+* MonoidAlgebra k G := { Finsupp.singleAddHom 1 with map_one' := rfl map_mul' := fun x y => by - -- Porting note (#10691): Was `rw`. simp only [ZeroHom.toFun_eq_coe, AddMonoidHom.toZeroHom_coe, singleAddHom_apply, single_mul_single, mul_one] } @@ -1367,7 +1366,6 @@ section Algebra def singleZeroRingHom [Semiring k] [AddMonoid G] : k →+* k[G] := { Finsupp.singleAddHom 0 with map_one' := rfl - -- Porting note (#10691): Was `rw`. map_mul' := fun x y => by simp only [singleAddHom, single_mul_single, zero_add] } /-- If two ring homomorphisms from `k[G]` are equal on all `single a 1` diff --git a/Mathlib/Algebra/Order/ToIntervalMod.lean b/Mathlib/Algebra/Order/ToIntervalMod.lean index 260764b5fe69c..0080b14e43fd5 100644 --- a/Mathlib/Algebra/Order/ToIntervalMod.lean +++ b/Mathlib/Algebra/Order/ToIntervalMod.lean @@ -760,7 +760,7 @@ private theorem toIxxMod_total' (a b c : α) : Thus if a ≠ b and b ≠ c then ({a-b} + {b-c}) + ({c-b} + {b-a}) = 2 * period, so one of `{a-b} + {b-c}` and `{c-b} + {b-a}` must be `≤ period` -/ have := congr_arg₂ (· + ·) (toIcoMod_add_toIocMod_zero hp a b) (toIcoMod_add_toIocMod_zero hp c b) - simp only [add_add_add_comm] at this -- Porting note (#10691): Was `rw` + simp only [add_add_add_comm] at this rw [_root_.add_comm (toIocMod _ _ _), add_add_add_comm, ← two_nsmul] at this replace := min_le_of_add_le_two_nsmul this.le rw [min_le_iff] at this diff --git a/Mathlib/Algebra/Polynomial/Laurent.lean b/Mathlib/Algebra/Polynomial/Laurent.lean index c139a6417371f..949d81ecba031 100644 --- a/Mathlib/Algebra/Polynomial/Laurent.lean +++ b/Mathlib/Algebra/Polynomial/Laurent.lean @@ -297,7 +297,7 @@ def trunc : R[T;T⁻¹] →+ R[X] := theorem trunc_C_mul_T (n : ℤ) (r : R) : trunc (C r * T n) = ite (0 ≤ n) (monomial n.toNat r) 0 := by apply (toFinsuppIso R).injective rw [← single_eq_C_mul_T, trunc, AddMonoidHom.coe_comp, Function.comp_apply] - -- Porting note (#10691): was `rw` + -- Porting note (#11224): was `rw` erw [comapDomain.addMonoidHom_apply Int.ofNat_injective] rw [toFinsuppIso_apply] -- Porting note: rewrote proof below relative to mathlib3. diff --git a/Mathlib/AlgebraicGeometry/OpenImmersion.lean b/Mathlib/AlgebraicGeometry/OpenImmersion.lean index 75394dff9384f..7b3769c508404 100644 --- a/Mathlib/AlgebraicGeometry/OpenImmersion.lean +++ b/Mathlib/AlgebraicGeometry/OpenImmersion.lean @@ -488,7 +488,7 @@ theorem range_pullback_snd_of_left : rw [← show _ = (pullback.snd f g).base from PreservesPullback.iso_hom_snd Scheme.forgetToTop f g, TopCat.coe_comp, Set.range_comp, Set.range_iff_surjective.mpr, ← @Set.preimage_univ _ _ (pullback.fst f.base g.base)] - -- Porting note (#10691): was `rw` + -- Porting note (#11224): was `rw` · erw [TopCat.pullback_snd_image_fst_preimage] rw [Set.image_univ] rfl @@ -505,7 +505,7 @@ theorem range_pullback_fst_of_right : rw [← show _ = (pullback.fst g f).base from PreservesPullback.iso_hom_fst Scheme.forgetToTop g f, TopCat.coe_comp, Set.range_comp, Set.range_iff_surjective.mpr, ← @Set.preimage_univ _ _ (pullback.snd g.base f.base)] - -- Porting note (#10691): was `rw` + -- Porting note (#11224): was `rw` · erw [TopCat.pullback_fst_image_snd_preimage] rw [Set.image_univ] rfl diff --git a/Mathlib/Data/DFinsupp/Interval.lean b/Mathlib/Data/DFinsupp/Interval.lean index 8588c90c87525..73f9fad8750ec 100644 --- a/Mathlib/Data/DFinsupp/Interval.lean +++ b/Mathlib/Data/DFinsupp/Interval.lean @@ -102,7 +102,6 @@ def rangeIcc (f g : Π₀ i, α i) : Π₀ i, Finset (α i) where (Multiset.not_mem_mono (Multiset.Le.subset <| Multiset.le_add_right _ _) h) have hg : g i = 0 := (gs.prop i).resolve_left (Multiset.not_mem_mono (Multiset.Le.subset <| Multiset.le_add_left _ _) h) - -- Porting note: was rw, but was rewriting under lambda, so changed to simp_rw simp_rw [hf, hg] exact Icc_self _⟩ diff --git a/Mathlib/LinearAlgebra/Alternating/DomCoprod.lean b/Mathlib/LinearAlgebra/Alternating/DomCoprod.lean index 847afe49e999e..001f2f17ef8d5 100644 --- a/Mathlib/LinearAlgebra/Alternating/DomCoprod.lean +++ b/Mathlib/LinearAlgebra/Alternating/DomCoprod.lean @@ -57,7 +57,7 @@ def domCoprod.summand (a : Mᵢ [⋀^ιa]→ₗ[R'] N₁) (b : Mᵢ [⋀^ιb]→ TensorProduct.smul_tmul'] simp only [Sum.map_inr, Perm.sumCongrHom_apply, Perm.sumCongr_apply, Sum.map_inl, Function.comp_apply, Perm.coe_mul] - -- Porting note (#10691): was `rw`. + -- Porting note (#11224): was `rw`. erw [← a.map_congr_perm fun i => v (σ₁ _), ← b.map_congr_perm fun i => v (σ₁ _)] theorem domCoprod.summand_mk'' (a : Mᵢ [⋀^ιa]→ₗ[R'] N₁) (b : Mᵢ [⋀^ιb]→ₗ[R'] N₂) diff --git a/Mathlib/LinearAlgebra/CliffordAlgebra/EvenEquiv.lean b/Mathlib/LinearAlgebra/CliffordAlgebra/EvenEquiv.lean index d699e29ff8b2d..119f96a21ae0d 100644 --- a/Mathlib/LinearAlgebra/CliffordAlgebra/EvenEquiv.lean +++ b/Mathlib/LinearAlgebra/CliffordAlgebra/EvenEquiv.lean @@ -117,7 +117,7 @@ def toEven : CliffordAlgebra Q →ₐ[R] CliffordAlgebra.even (Q' Q) := by theorem toEven_ι (m : M) : (toEven Q (ι Q m) : CliffordAlgebra (Q' Q)) = e0 Q * v Q m := by rw [toEven, CliffordAlgebra.lift_ι_apply] - -- Porting note (#10691): was `rw` + -- Porting note (#11224): was `rw` erw [LinearMap.codRestrict_apply] rw [LinearMap.coe_comp, Function.comp_apply, LinearMap.mulLeft_apply] diff --git a/Mathlib/Probability/Kernel/MeasurableIntegral.lean b/Mathlib/Probability/Kernel/MeasurableIntegral.lean index dc0acc26cac82..b35f6f0a1fcab 100644 --- a/Mathlib/Probability/Kernel/MeasurableIntegral.lean +++ b/Mathlib/Probability/Kernel/MeasurableIntegral.lean @@ -166,7 +166,7 @@ theorem _root_.Measurable.lintegral_kernel_prod_right {f : α → β → ℝ≥0 (fun a => ∫⁻ b, g₁ (a, b) ∂κ a) + fun a => ∫⁻ b, g₂ (a, b) ∂κ a := by ext1 a rw [Pi.add_apply] - -- Porting note (#10691): was `rw` (`Function.comp` reducibility) + -- Porting note (#11224): was `rw` (`Function.comp` reducibility) erw [lintegral_add_left (g₁.measurable.comp measurable_prod_mk_left)] simp_rw [Function.comp_apply] rw [h_add] diff --git a/Mathlib/RingTheory/Localization/Ideal.lean b/Mathlib/RingTheory/Localization/Ideal.lean index 990c36d781bd6..dfcf60e25dc37 100644 --- a/Mathlib/RingTheory/Localization/Ideal.lean +++ b/Mathlib/RingTheory/Localization/Ideal.lean @@ -201,9 +201,7 @@ theorem surjective_quotientMap_of_maximal_of_localization {I : Ideal S} [I.IsPri (Ideal.Quotient.eq_zero_iff_mem.2 (Ideal.mem_comap.2 (Ideal.Quotient.eq_zero_iff_mem.1 hn)))) (_root_.trans hn ?_)) - -- Porting note (#10691): was `rw`, but this took extremely long. - refine Eq.trans ?_ (RingHom.map_mul (Ideal.Quotient.mk I) (algebraMap R S m) (mk' S 1 ⟨m, hm⟩)) - rw [← mk'_eq_mul_mk'_one, mk'_self, RingHom.map_one] + rw [← map_mul, ← mk'_eq_mul_mk'_one, mk'_self, RingHom.map_one] open nonZeroDivisors From d951eef6bb39b3b78cbd1fba9c515c3770a9d20a Mon Sep 17 00:00:00 2001 From: Christian Merten Date: Thu, 17 Oct 2024 10:54:51 +0000 Subject: [PATCH 069/505] feat(AlgebraicGeometry): stability properties of separated (#17861) Separated is multiplicative, stable under base change and local at the target. --- .../Morphisms/Separated.lean | 21 ++++++++++++++++--- 1 file changed, 18 insertions(+), 3 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Separated.lean b/Mathlib/AlgebraicGeometry/Morphisms/Separated.lean index 5831774a682e8..369855b7c262d 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Separated.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Separated.lean @@ -16,9 +16,6 @@ A morphism of schemes is separated if its diagonal morphism is a closed immmersi ## TODO -- Show separated is stable under composition and base change (this is immediate if - we show that closed immersions are stable under base change). -- Show separated is local at the target. - Show affine morphisms are separated. -/ @@ -60,6 +57,24 @@ instance : MorphismProperty.RespectsIso @IsSeparated := by instance (priority := 900) [IsSeparated f] : QuasiSeparated f where +instance stableUnderComposition : MorphismProperty.IsStableUnderComposition @IsSeparated := by + rw [isSeparated_eq_diagonal_isClosedImmersion] + exact MorphismProperty.diagonal_isStableUnderComposition IsClosedImmersion.stableUnderBaseChange + +instance {Z : Scheme.{u}} (g : Y ⟶ Z) [IsSeparated f] [IsSeparated g] : IsSeparated (f ≫ g) := + stableUnderComposition.comp_mem f g inferInstance inferInstance + +instance : MorphismProperty.IsMultiplicative @IsSeparated where + id_mem _ := inferInstance + +lemma stableUnderBaseChange : MorphismProperty.StableUnderBaseChange @IsSeparated := by + rw [isSeparated_eq_diagonal_isClosedImmersion] + exact MorphismProperty.StableUnderBaseChange.diagonal IsClosedImmersion.stableUnderBaseChange + +instance : IsLocalAtTarget @IsSeparated := by + rw [isSeparated_eq_diagonal_isClosedImmersion] + infer_instance + end IsSeparated end AlgebraicGeometry From d51a8c9ac5f7970ce1d50207a52ff800aa95d956 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Thu, 17 Oct 2024 11:30:00 +0000 Subject: [PATCH 070/505] feat(LinearAlgebra/TensorProduct/Quotient): add some isomorphisms (#17803) Moves: - quotTensorEquivQuotSMul -> TensorProduct.quotTensorEquivQuotSMul - tensorQuotEquivQuotSMul -> TensorProduct.tensorQuotEquivQuotSMul --- Mathlib.lean | 1 + .../LinearAlgebra/TensorProduct/Quotient.lean | 241 ++++++++++++++++++ .../TensorProduct/RightExactness.lean | 79 ------ Mathlib/RingTheory/QuotSMulTop.lean | 1 + 4 files changed, 243 insertions(+), 79 deletions(-) create mode 100644 Mathlib/LinearAlgebra/TensorProduct/Quotient.lean diff --git a/Mathlib.lean b/Mathlib.lean index 9005c78e9dcb7..0439d6b8d5e5f 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3173,6 +3173,7 @@ import Mathlib.LinearAlgebra.TensorProduct.Matrix import Mathlib.LinearAlgebra.TensorProduct.Opposite import Mathlib.LinearAlgebra.TensorProduct.Pi import Mathlib.LinearAlgebra.TensorProduct.Prod +import Mathlib.LinearAlgebra.TensorProduct.Quotient import Mathlib.LinearAlgebra.TensorProduct.RightExactness import Mathlib.LinearAlgebra.TensorProduct.Subalgebra import Mathlib.LinearAlgebra.TensorProduct.Submodule diff --git a/Mathlib/LinearAlgebra/TensorProduct/Quotient.lean b/Mathlib/LinearAlgebra/TensorProduct/Quotient.lean new file mode 100644 index 0000000000000..f3cea863abf53 --- /dev/null +++ b/Mathlib/LinearAlgebra/TensorProduct/Quotient.lean @@ -0,0 +1,241 @@ +/- +Copyright (c) 2024 Antoine Chambert-Loir. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Antoine Chambert-Loir, Jujian Zhang +-/ + +import Mathlib.RingTheory.Ideal.Operations +import Mathlib.LinearAlgebra.TensorProduct.Basic +import Mathlib.LinearAlgebra.Quotient +import Mathlib.LinearAlgebra.Prod +import Mathlib.RingTheory.Ideal.Quotient + +/-! + +# Interaction between Quotients and Tensor Products + +This file contains constructions that relate quotients and tensor products. +Let `M, N` be `R`-modules, `m ≤ M` and `n ≤ N` be an `R`-submodules and `I ≤ R` an ideal. We prove +the following isomorphisms: + +## Main results +- `TensorProduct.quotientTensorQuotientEquiv`: + `(M ⧸ m) ⊗[R] (N ⧸ n) ≃ₗ[R] (M ⊗[R] N) ⧸ (m ⊗ N ⊔ M ⊗ n)` +- `TensorProduct.quotientTensorEquiv`: + `(M ⧸ m) ⊗[R] N ≃ₗ[R] (M ⊗[R] N) ⧸ (m ⊗ N)` +- `TensorProduct.tensorQuotientEquiv`: + `M ⊗[R] (N ⧸ n) ≃ₗ[R] (M ⊗[R] N) ⧸ (M ⊗ n)` +- `TensorProduct.quotTensorEquivQuotSMul`: + `(R ⧸ I) ⊗[R] M ≃ₗ[R] M ⧸ (I • M)` +- `TensorProduct.tensorQuotEquivQuotSMul`: + `M ⊗[R] (R ⧸ I) ≃ₗ[R] M ⧸ (I • M)` + +## Tags + +Quotient, Tensor Product + +-/ + +namespace TensorProduct + +variable {R M N : Type*} [CommRing R] +variable [AddCommGroup M] [Module R M] [AddCommGroup N] [Module R N] + +attribute [local ext high] ext LinearMap.prod_ext + +/-- +Let `M, N` be `R`-modules, `m ≤ M` and `n ≤ N` be an `R`-submodules. Then we have a linear +isomorphism between tensor products of the quotients and the quotient of the tensor product: +`(M ⧸ m) ⊗[R] (N ⧸ n) ≃ₗ[R] (M ⊗[R] N) ⧸ (m ⊗ N ⊔ M ⊗ n)`. +-/ +noncomputable def quotientTensorQuotientEquiv (m : Submodule R M) (n : Submodule R N) : + (M ⧸ (m : Submodule R M)) ⊗[R] (N ⧸ (n : Submodule R N)) ≃ₗ[R] + (M ⊗[R] N) ⧸ + (LinearMap.range (map m.subtype LinearMap.id) ⊔ + LinearMap.range (map LinearMap.id n.subtype)) := + LinearEquiv.ofLinear + (lift <| Submodule.liftQ _ (LinearMap.flip <| Submodule.liftQ _ + ((mk R (M := M) (N := N)).flip.compr₂ (Submodule.mkQ _)) fun x hx => by + ext y + simp only [LinearMap.compr₂_apply, LinearMap.flip_apply, mk_apply, Submodule.mkQ_apply, + LinearMap.zero_apply, Submodule.Quotient.mk_eq_zero] + exact Submodule.mem_sup_right ⟨y ⊗ₜ ⟨x, hx⟩, rfl⟩) fun x hx => by + ext y + simp only [LinearMap.coe_comp, Function.comp_apply, Submodule.mkQ_apply, LinearMap.flip_apply, + Submodule.liftQ_apply, LinearMap.compr₂_apply, mk_apply, LinearMap.zero_comp, + LinearMap.zero_apply, Submodule.Quotient.mk_eq_zero] + exact Submodule.mem_sup_left ⟨⟨x, hx⟩ ⊗ₜ y, rfl⟩) + (Submodule.liftQ _ (map (Submodule.mkQ _) (Submodule.mkQ _)) fun x hx => by + rw [Submodule.mem_sup] at hx + rcases hx with ⟨_, ⟨a, rfl⟩, _, ⟨b, rfl⟩, rfl⟩ + simp only [LinearMap.mem_ker, map_add] + set f := (map m.mkQ n.mkQ) ∘ₗ (map m.subtype LinearMap.id) + set g := (map m.mkQ n.mkQ) ∘ₗ (map LinearMap.id n.subtype) + have eq : LinearMap.coprod f g = 0 := by + ext x y + · simp [f, Submodule.Quotient.mk_eq_zero _ |>.2 x.2] + · simp [g, Submodule.Quotient.mk_eq_zero _ |>.2 y.2] + exact congr($eq (a, b))) + (by ext; simp) (by ext; simp) + +@[simp] +lemma quotientTensorQuotientEquiv_apply_tmul_mk_tmul_mk + (m : Submodule R M) (n : Submodule R N) (x : M) (y : N) : + quotientTensorQuotientEquiv m n + (Submodule.Quotient.mk x ⊗ₜ[R] Submodule.Quotient.mk y) = + Submodule.Quotient.mk (x ⊗ₜ y) := rfl + +@[simp] +lemma quotientTensorQuotientEquiv_symm_apply_mk_tmul + (m : Submodule R M) (n : Submodule R N) (x : M) (y : N) : + (quotientTensorQuotientEquiv m n).symm (Submodule.Quotient.mk (x ⊗ₜ y)) = + Submodule.Quotient.mk x ⊗ₜ[R] Submodule.Quotient.mk y := rfl + +variable (N) in +/-- +Let `M, N` be `R`-modules, `m ≤ M` be an `R`-submodule. Then we have a linear isomorphism between +tensor products of the quotient and the quotient of the tensor product: +`(M ⧸ m) ⊗[R] N ≃ₗ[R] (M ⊗[R] N) ⧸ (m ⊗ N)`. +-/ +noncomputable def quotientTensorEquiv (m : Submodule R M) : + (M ⧸ (m : Submodule R M)) ⊗[R] N ≃ₗ[R] + (M ⊗[R] N) ⧸ (LinearMap.range (map m.subtype (LinearMap.id : N →ₗ[R] N))) := + congr (LinearEquiv.refl _ _) ((Submodule.quotEquivOfEqBot _ rfl).symm) ≪≫ₗ + quotientTensorQuotientEquiv (N := N) m ⊥ ≪≫ₗ + Submodule.Quotient.equiv _ _ (LinearEquiv.refl _ _) (by + simp only [Submodule.map_sup] + erw [Submodule.map_id, Submodule.map_id] + simp only [sup_eq_left] + rw [map_range_eq_span_tmul, map_range_eq_span_tmul] + aesop) + +@[simp] +lemma quotientTensorEquiv_apply_tmul_mk (m : Submodule R M) (x : M) (y : N) : + quotientTensorEquiv N m (Submodule.Quotient.mk x ⊗ₜ[R] y) = + Submodule.Quotient.mk (x ⊗ₜ y) := + rfl + +@[simp] +lemma quotientTensorEquiv_symm_apply_mk_tmul (m : Submodule R M) (x : M) (y : N) : + (quotientTensorEquiv N m).symm (Submodule.Quotient.mk (x ⊗ₜ y)) = + Submodule.Quotient.mk x ⊗ₜ[R] y := + rfl + +variable (M) in +/-- +Let `M, N` be `R`-modules, `n ≤ N` be an `R`-submodule. Then we have a linear isomorphism between +tensor products of the quotient and the quotient of the tensor product: +`M ⊗[R] (N ⧸ n) ≃ₗ[R] (M ⊗[R] N) ⧸ (M ⊗ n)`. +-/ +noncomputable def tensorQuotientEquiv (n : Submodule R N) : + M ⊗[R] (N ⧸ (n : Submodule R N)) ≃ₗ[R] + (M ⊗[R] N) ⧸ (LinearMap.range (map (LinearMap.id : M →ₗ[R] M) n.subtype)) := + congr ((Submodule.quotEquivOfEqBot _ rfl).symm) (LinearEquiv.refl _ _) ≪≫ₗ + quotientTensorQuotientEquiv (⊥ : Submodule R M) n ≪≫ₗ + Submodule.Quotient.equiv _ _ (LinearEquiv.refl _ _) (by + simp only [Submodule.map_sup] + erw [Submodule.map_id, Submodule.map_id] + simp only [sup_eq_right] + rw [map_range_eq_span_tmul, map_range_eq_span_tmul] + aesop) + +@[simp] +lemma tensorQuotientEquiv_apply_mk_tmul (n : Submodule R N) (x : M) (y : N) : + tensorQuotientEquiv M n (x ⊗ₜ[R] Submodule.Quotient.mk y) = + Submodule.Quotient.mk (x ⊗ₜ y) := + rfl + +@[simp] +lemma tensorQuotientEquiv_symm_apply_tmul_mk (n : Submodule R N) (x : M) (y : N) : + (tensorQuotientEquiv M n).symm (Submodule.Quotient.mk (x ⊗ₜ y)) = + x ⊗ₜ[R] Submodule.Quotient.mk y := + rfl + +variable (M) in +/-- Left tensoring a module with a quotient of the ring is the same as +quotienting that module by the corresponding submodule. -/ +noncomputable def quotTensorEquivQuotSMul (I : Ideal R) : + ((R ⧸ I) ⊗[R] M) ≃ₗ[R] M ⧸ (I • (⊤ : Submodule R M)) := + quotientTensorEquiv M I ≪≫ₗ + Submodule.Quotient.equiv (M := R ⊗[R] M) (N := M) (f := TensorProduct.lid R M) (hf := rfl) ≪≫ₗ + Submodule.Quotient.equiv _ _ (LinearEquiv.refl R M) (by + erw [Submodule.map_id] + rw [TensorProduct.map_range_eq_span_tmul, Submodule.map_span] + refine le_antisymm (Submodule.span_le.2 ?_) (Submodule.map₂_le.2 ?_) + · rintro _ ⟨_, ⟨r, m, rfl⟩, rfl⟩ + simp only [Submodule.coe_subtype, LinearMap.id_coe, id_eq, lid_tmul, SetLike.mem_coe] + apply Submodule.apply_mem_map₂ <;> aesop + · rintro r hr m - + simp only [Submodule.coe_subtype, LinearMap.id_coe, id_eq, Subtype.exists, exists_prop, + LinearMap.lsmul_apply] + refine Submodule.subset_span ?_ + simp only [Set.mem_image, Set.mem_setOf_eq] + exact ⟨r ⊗ₜ m, ⟨r, hr, m, rfl⟩, rfl⟩) + +variable (M) in +/-- Right tensoring a module with a quotient of the ring is the same as +quotienting that module by the corresponding submodule. -/ +noncomputable def tensorQuotEquivQuotSMul (I : Ideal R) : + (M ⊗[R] (R ⧸ I)) ≃ₗ[R] M ⧸ (I • (⊤ : Submodule R M)) := + TensorProduct.comm _ _ _ ≪≫ₗ quotTensorEquivQuotSMul M I + +@[simp] +lemma quotTensorEquivQuotSMul_mk_tmul (I : Ideal R) (r : R) (x : M) : + quotTensorEquivQuotSMul M I (Ideal.Quotient.mk I r ⊗ₜ[R] x) = + Submodule.Quotient.mk (r • x) := + (quotTensorEquivQuotSMul M I).eq_symm_apply.mp <| + Eq.trans (congrArg (· ⊗ₜ[R] x) <| + Eq.trans (congrArg (Ideal.Quotient.mk I) + (Eq.trans (smul_eq_mul R) (mul_one r))).symm <| + Submodule.Quotient.mk_smul I r 1) <| + smul_tmul r _ x + +lemma quotTensorEquivQuotSMul_comp_mkQ_rTensor (I : Ideal R) : + quotTensorEquivQuotSMul M I ∘ₗ I.mkQ.rTensor M = + (I • ⊤ : Submodule R M).mkQ ∘ₗ TensorProduct.lid R M := + TensorProduct.ext' (quotTensorEquivQuotSMul_mk_tmul I) + +@[simp] +lemma quotTensorEquivQuotSMul_symm_mk (I : Ideal R) (x : M) : + (quotTensorEquivQuotSMul M I).symm (Submodule.Quotient.mk x) = 1 ⊗ₜ[R] x := + rfl + +lemma quotTensorEquivQuotSMul_symm_comp_mkQ (I : Ideal R) : + (quotTensorEquivQuotSMul M I).symm ∘ₗ (I • ⊤ : Submodule R M).mkQ = + TensorProduct.mk R (R ⧸ I) M 1 := + LinearMap.ext (quotTensorEquivQuotSMul_symm_mk I) + +lemma quotTensorEquivQuotSMul_comp_mk (I : Ideal R) : + quotTensorEquivQuotSMul M I ∘ₗ TensorProduct.mk R (R ⧸ I) M 1 = + Submodule.mkQ (I • ⊤) := + Eq.symm <| (LinearEquiv.toLinearMap_symm_comp_eq _ _).mp <| + quotTensorEquivQuotSMul_symm_comp_mkQ I + +@[simp] +lemma tensorQuotEquivQuotSMul_tmul_mk (I : Ideal R) (x : M) (r : R) : + tensorQuotEquivQuotSMul M I (x ⊗ₜ[R] Ideal.Quotient.mk I r) = + Submodule.Quotient.mk (r • x) := + quotTensorEquivQuotSMul_mk_tmul I r x + +lemma tensorQuotEquivQuotSMul_comp_mkQ_lTensor (I : Ideal R) : + tensorQuotEquivQuotSMul M I ∘ₗ I.mkQ.lTensor M = + (I • ⊤ : Submodule R M).mkQ ∘ₗ TensorProduct.rid R M := + TensorProduct.ext' (tensorQuotEquivQuotSMul_tmul_mk I) + +@[simp] +lemma tensorQuotEquivQuotSMul_symm_mk (I : Ideal R) (x : M) : + (tensorQuotEquivQuotSMul M I).symm (Submodule.Quotient.mk x) = x ⊗ₜ[R] 1 := + rfl + +lemma tensorQuotEquivQuotSMul_symm_comp_mkQ (I : Ideal R) : + (tensorQuotEquivQuotSMul M I).symm ∘ₗ (I • ⊤ : Submodule R M).mkQ = + (TensorProduct.mk R M (R ⧸ I)).flip 1 := + LinearMap.ext (tensorQuotEquivQuotSMul_symm_mk I) + +lemma tensorQuotEquivQuotSMul_comp_mk (I : Ideal R) : + tensorQuotEquivQuotSMul M I ∘ₗ (TensorProduct.mk R M (R ⧸ I)).flip 1 = + Submodule.mkQ (I • ⊤) := + Eq.symm <| (LinearEquiv.toLinearMap_symm_comp_eq _ _).mp <| + tensorQuotEquivQuotSMul_symm_comp_mkQ I + +end TensorProduct diff --git a/Mathlib/LinearAlgebra/TensorProduct/RightExactness.lean b/Mathlib/LinearAlgebra/TensorProduct/RightExactness.lean index c55c56caf1c8e..a61f53baf761c 100644 --- a/Mathlib/LinearAlgebra/TensorProduct/RightExactness.lean +++ b/Mathlib/LinearAlgebra/TensorProduct/RightExactness.lean @@ -6,7 +6,6 @@ Authors: Antoine Chambert-Loir import Mathlib.Algebra.Exact import Mathlib.RingTheory.TensorProduct.Basic -import Mathlib.RingTheory.Ideal.Quotient /-! # Right-exactness properties of tensor product @@ -423,84 +422,6 @@ theorem TensorProduct.mk_surjective (S) [Semiring S] [Algebra R S] rw [Algebra.algebraMap_eq_smul_one, smul_tmul] exact ⟨x • y, rfl⟩ -/-- Left tensoring a module with a quotient of the ring is the same as -quotienting that module by the corresponding submodule. -/ -noncomputable def quotTensorEquivQuotSMul (I : Ideal R) : - (R ⧸ I) ⊗[R] M ≃ₗ[R] M ⧸ (I • ⊤ : Submodule R M) := - (rTensor.equiv M (exact_subtype_mkQ I) I.mkQ_surjective).symm.trans <| - Submodule.Quotient.equiv _ _ (TensorProduct.lid R M) <| - Eq.trans (LinearMap.range_comp _ _).symm <| - Eq.trans ((Submodule.topEquiv.lTensor I).range_comp _).symm <| - Eq.symm <| Eq.trans (map₂_eq_range_lift_comp_mapIncl _ _ _) <| - congrArg _ (TensorProduct.ext' (fun _ _ => rfl)) - -/-- Right tensoring a module with a quotient of the ring is the same as -quotienting that module by the corresponding submodule. -/ -noncomputable def tensorQuotEquivQuotSMul (I : Ideal R) : - M ⊗[R] (R ⧸ I) ≃ₗ[R] M ⧸ (I • ⊤ : Submodule R M) := - TensorProduct.comm R M _ ≪≫ₗ quotTensorEquivQuotSMul M I - -variable {M} - -@[simp] -lemma quotTensorEquivQuotSMul_mk_tmul (I : Ideal R) (r : R) (x : M) : - quotTensorEquivQuotSMul M I (Ideal.Quotient.mk I r ⊗ₜ[R] x) = - Submodule.Quotient.mk (r • x) := - (quotTensorEquivQuotSMul M I).eq_symm_apply.mp <| - Eq.trans (congrArg (· ⊗ₜ[R] x) <| - Eq.trans (congrArg (Ideal.Quotient.mk I) - (Eq.trans (smul_eq_mul R) (mul_one r))).symm <| - Submodule.Quotient.mk_smul I r 1) <| - smul_tmul r _ x - -lemma quotTensorEquivQuotSMul_comp_mkQ_rTensor (I : Ideal R) : - quotTensorEquivQuotSMul M I ∘ₗ I.mkQ.rTensor M = - (I • ⊤ : Submodule R M).mkQ ∘ₗ TensorProduct.lid R M := - TensorProduct.ext' (quotTensorEquivQuotSMul_mk_tmul I) - -@[simp] -lemma quotTensorEquivQuotSMul_symm_mk (I : Ideal R) (x : M) : - (quotTensorEquivQuotSMul M I).symm (Submodule.Quotient.mk x) = 1 ⊗ₜ[R] x := - rfl - -lemma quotTensorEquivQuotSMul_symm_comp_mkQ (I : Ideal R) : - (quotTensorEquivQuotSMul M I).symm ∘ₗ (I • ⊤ : Submodule R M).mkQ = - TensorProduct.mk R (R ⧸ I) M 1 := - LinearMap.ext (quotTensorEquivQuotSMul_symm_mk I) - -lemma quotTensorEquivQuotSMul_comp_mk (I : Ideal R) : - quotTensorEquivQuotSMul M I ∘ₗ TensorProduct.mk R (R ⧸ I) M 1 = - Submodule.mkQ (I • ⊤) := - Eq.symm <| (LinearEquiv.toLinearMap_symm_comp_eq _ _).mp <| - quotTensorEquivQuotSMul_symm_comp_mkQ I - -@[simp] -lemma tensorQuotEquivQuotSMul_tmul_mk (I : Ideal R) (x : M) (r : R) : - tensorQuotEquivQuotSMul M I (x ⊗ₜ[R] Ideal.Quotient.mk I r) = - Submodule.Quotient.mk (r • x) := - quotTensorEquivQuotSMul_mk_tmul I r x - -lemma tensorQuotEquivQuotSMul_comp_mkQ_lTensor (I : Ideal R) : - tensorQuotEquivQuotSMul M I ∘ₗ I.mkQ.lTensor M = - (I • ⊤ : Submodule R M).mkQ ∘ₗ TensorProduct.rid R M := - TensorProduct.ext' (tensorQuotEquivQuotSMul_tmul_mk I) - -@[simp] -lemma tensorQuotEquivQuotSMul_symm_mk (I : Ideal R) (x : M) : - (tensorQuotEquivQuotSMul M I).symm (Submodule.Quotient.mk x) = x ⊗ₜ[R] 1 := - rfl - -lemma tensorQuotEquivQuotSMul_symm_comp_mkQ (I : Ideal R) : - (tensorQuotEquivQuotSMul M I).symm ∘ₗ (I • ⊤ : Submodule R M).mkQ = - (TensorProduct.mk R M (R ⧸ I)).flip 1 := - LinearMap.ext (tensorQuotEquivQuotSMul_symm_mk I) - -lemma tensorQuotEquivQuotSMul_comp_mk (I : Ideal R) : - tensorQuotEquivQuotSMul M I ∘ₗ (TensorProduct.mk R M (R ⧸ I)).flip 1 = - Submodule.mkQ (I • ⊤) := - Eq.symm <| (LinearEquiv.toLinearMap_symm_comp_eq _ _).mp <| - tensorQuotEquivQuotSMul_symm_comp_mkQ I - end Modules section Algebras diff --git a/Mathlib/RingTheory/QuotSMulTop.lean b/Mathlib/RingTheory/QuotSMulTop.lean index 50ffb52e03183..2b2dfe48be108 100644 --- a/Mathlib/RingTheory/QuotSMulTop.lean +++ b/Mathlib/RingTheory/QuotSMulTop.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Brendan Murphy -/ import Mathlib.LinearAlgebra.TensorProduct.RightExactness +import Mathlib.LinearAlgebra.TensorProduct.Quotient /-! # Reducing a module modulo an element of the ring From 7d50a88f4d70f386ac24e982cf686583e96e2959 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Thu, 17 Oct 2024 11:36:58 +0000 Subject: [PATCH 071/505] feat(Algebra): integralClosure as IntermediateField (#14206) + Add lemmas about integrality and membership in an subalgebra of inverse elements. + As a consequence, show the integral closure of a field in a division ring is an intermediate field `IntermediateField.algClosure`. + Golf some lemmas along the way. Co-authored-by: Junyan Xu --- Mathlib/FieldTheory/Adjoin.lean | 55 +++++-------- .../IntermediateField/Algebraic.lean | 24 +++++- .../FieldTheory/IntermediateField/Basic.lean | 30 ++----- Mathlib/FieldTheory/PurelyInseparable.lean | 2 +- .../LinearAlgebra/FiniteDimensional/Defs.lean | 7 +- Mathlib/RingTheory/Algebraic.lean | 2 +- .../IntegralClosure/Algebra/Basic.lean | 17 ++-- .../IsIntegralClosure/Basic.lean | 82 ++++++++++++++++--- 8 files changed, 137 insertions(+), 82 deletions(-) diff --git a/Mathlib/FieldTheory/Adjoin.lean b/Mathlib/FieldTheory/Adjoin.lean index 3fa951422ab8a..aa48dc241889b 100644 --- a/Mathlib/FieldTheory/Adjoin.lean +++ b/Mathlib/FieldTheory/Adjoin.lean @@ -70,7 +70,7 @@ section Lattice variable {F : Type*} [Field F] {E : Type*} [Field E] [Algebra F E] @[simp] -theorem adjoin_le_iff {S : Set E} {T : IntermediateField F E} : adjoin F S ≤ T ↔ S ≤ T := +theorem adjoin_le_iff {S : Set E} {T : IntermediateField F E} : adjoin F S ≤ T ↔ S ⊆ T := ⟨fun H => le_trans (le_trans Set.subset_union_right Subfield.subset_closure) H, fun H => (@Subfield.closure_le E _ (Set.range (algebraMap F E) ∪ S) T.toSubfield).mpr (Set.union_subset (IntermediateField.set_range_subset T) H)⟩ @@ -263,7 +263,7 @@ instance isScalarTower_over_bot : IsScalarTower (⊥ : IntermediateField F E) F This is the intermediate field version of `Subalgebra.topEquiv`. -/ @[simps!] def topEquiv : (⊤ : IntermediateField F E) ≃ₐ[F] E := - (Subalgebra.equivOfEq _ _ top_toSubalgebra).trans Subalgebra.topEquiv + Subalgebra.topEquiv -- Porting note: this theorem is now generated by the `@[simps!]` above. @@ -400,24 +400,14 @@ theorem adjoin_subset_adjoin_iff {F' : Type*} [Field F'] [Algebra F' E] {S S' : theorem adjoin_adjoin_left (T : Set E) : (adjoin (adjoin F S) T).restrictScalars _ = adjoin F (S ∪ T) := by rw [SetLike.ext'_iff] - change (↑(adjoin (adjoin F S) T) : Set E) = _ - apply Set.eq_of_subset_of_subset <;> rw [adjoin_subset_adjoin_iff] <;> constructor + change (adjoin (adjoin F S) T : Set E) = _ + apply subset_antisymm <;> rw [adjoin_subset_adjoin_iff] <;> constructor · rintro _ ⟨⟨x, hx⟩, rfl⟩; exact adjoin.mono _ _ _ Set.subset_union_left hx · exact subset_adjoin_of_subset_right _ _ Set.subset_union_right --- Porting note: original proof times out - · rintro x ⟨f, rfl⟩ - refine Subfield.subset_closure ?_ - left - exact ⟨f, rfl⟩ --- Porting note: original proof times out - · refine Set.union_subset (fun x hx => Subfield.subset_closure ?_) - (fun x hx => Subfield.subset_closure ?_) - · left - refine ⟨⟨x, Subfield.subset_closure ?_⟩, rfl⟩ - right - exact hx - · right - exact hx + · exact Set.range_subset_iff.mpr fun f ↦ Subfield.subset_closure (.inl ⟨f, rfl⟩) + · exact Set.union_subset + (fun x hx ↦ Subfield.subset_closure <| .inl ⟨⟨x, Subfield.subset_closure (.inr hx)⟩, rfl⟩) + (fun x hx ↦ Subfield.subset_closure <| .inr hx) @[simp] theorem adjoin_insert_adjoin (x : E) : @@ -435,14 +425,10 @@ theorem adjoin_adjoin_comm (T : Set E) : rw [adjoin_adjoin_left, adjoin_adjoin_left, Set.union_comm] theorem adjoin_map {E' : Type*} [Field E'] [Algebra F E'] (f : E →ₐ[F] E') : - (adjoin F S).map f = adjoin F (f '' S) := by - ext x - show - x ∈ (Subfield.closure (Set.range (algebraMap F E) ∪ S)).map (f : E →+* E') ↔ - x ∈ Subfield.closure (Set.range (algebraMap F E') ∪ f '' S) - rw [RingHom.map_field_closure, Set.image_union, ← Set.range_comp, ← RingHom.coe_comp, - f.comp_algebraMap] - rfl + (adjoin F S).map f = adjoin F (f '' S) := + le_antisymm + (map_le_iff_le_comap.mpr <| adjoin_le_iff.mpr fun x hx ↦ subset_adjoin _ _ ⟨x, hx, rfl⟩) + (adjoin_le_iff.mpr <| Set.monotone_image <| subset_adjoin _ _) @[simp] theorem lift_adjoin (K : IntermediateField F E) (S : Set K) : @@ -510,9 +496,9 @@ theorem eq_adjoin_of_eq_algebra_adjoin (K : IntermediateField F E) (h : K.toSubalgebra = Algebra.adjoin F S) : K = adjoin F S := by apply toSubalgebra_injective rw [h] - refine (adjoin_eq_algebra_adjoin F _ ?_).symm - intro x - convert K.inv_mem (x := x) <;> rw [← h] <;> rfl + refine (adjoin_eq_algebra_adjoin F _ fun x ↦ ?_).symm + rw [← h] + exact K.inv_mem theorem adjoin_eq_top_of_algebra (hS : Algebra.adjoin F S = ⊤) : adjoin F S = ⊤ := top_le_iff.mp (hS.symm.trans_le <| algebra_adjoin_le_adjoin F S) @@ -600,12 +586,9 @@ theorem adjoin_simple_comm (β : E) : F⟮α⟯⟮β⟯.restrictScalars F = F⟮ variable {F} {α} theorem adjoin_algebraic_toSubalgebra {S : Set E} (hS : ∀ x ∈ S, IsAlgebraic F x) : - (IntermediateField.adjoin F S).toSubalgebra = Algebra.adjoin F S := by - simp only [isAlgebraic_iff_isIntegral] at hS - have : Algebra.IsIntegral F (Algebra.adjoin F S) := by - rwa [← le_integralClosure_iff_isIntegral, Algebra.adjoin_le_iff] - have : IsField (Algebra.adjoin F S) := isField_of_isIntegral_of_isField' (Field.toIsField F) - rw [← ((Algebra.adjoin F S).toIntermediateField' this).eq_adjoin_of_eq_algebra_adjoin F S] <;> rfl + (IntermediateField.adjoin F S).toSubalgebra = Algebra.adjoin F S := + adjoin_eq_algebra_adjoin _ _ fun _ ↦ + (Algebra.IsIntegral.adjoin fun x hx ↦ (hS x hx).isIntegral).inv_mem theorem adjoin_simple_toSubalgebra_of_integral (hα : IsIntegral F α) : F⟮α⟯.toSubalgebra = Algebra.adjoin F {α} := by @@ -884,7 +867,7 @@ variable {F : Type*} [Field F] {E : Type*} [Field E] [Algebra F E] {α : E} {S : @[simp] theorem adjoin_eq_bot_iff : adjoin F S = ⊥ ↔ S ⊆ (⊥ : IntermediateField F E) := by - rw [eq_bot_iff, adjoin_le_iff]; rfl + rw [eq_bot_iff, adjoin_le_iff] /- Porting note: this was tagged `simp`. -/ theorem adjoin_simple_eq_bot_iff : F⟮α⟯ = ⊥ ↔ α ∈ (⊥ : IntermediateField F E) := by diff --git a/Mathlib/FieldTheory/IntermediateField/Algebraic.lean b/Mathlib/FieldTheory/IntermediateField/Algebraic.lean index 55ecc96148306..7769f62c1eccb 100644 --- a/Mathlib/FieldTheory/IntermediateField/Algebraic.lean +++ b/Mathlib/FieldTheory/IntermediateField/Algebraic.lean @@ -18,6 +18,26 @@ open Module variable {K : Type*} {L : Type*} [Field K] [Field L] [Algebra K L] {S : IntermediateField K L} +theorem IntermediateField.coe_isIntegral_iff {R : Type*} [CommRing R] [Algebra R K] [Algebra R L] + [IsScalarTower R K L] {x : S} : IsIntegral R (x : L) ↔ IsIntegral R x := + isIntegral_algHom_iff (S.val.restrictScalars R) Subtype.val_injective + +/-- Turn an algebraic subalgebra into an intermediate field, `Subalgebra.IsAlgebraic` version. -/ +def Subalgebra.IsAlgebraic.toIntermediateField {S : Subalgebra K L} (hS : S.IsAlgebraic) : + IntermediateField K L where + toSubalgebra := S + inv_mem' x hx := Algebra.adjoin_le_iff.mpr + (Set.singleton_subset_iff.mpr hx) (hS x hx).isIntegral.inv_mem_adjoin + +/-- Turn an algebraic subalgebra into an intermediate field, `Algebra.IsAlgebraic` version. -/ +abbrev Algebra.IsAlgebraic.toIntermediateField (S : Subalgebra K L) [Algebra.IsAlgebraic K S] : + IntermediateField K L := (S.isAlgebraic_iff.mpr ‹_›).toIntermediateField + +/-- The algebraic closure of a field `K` in an extension `L`, the subalgebra `integralClosure K L` +upgraded to an intermediate field (when `K` and `L` are both fields). -/ +def algebraicClosure : IntermediateField K L := + Algebra.IsAlgebraic.toIntermediateField (integralClosure K L) + namespace IntermediateField section FiniteDimensional @@ -86,8 +106,8 @@ end FiniteDimensional theorem isAlgebraic_iff {x : S} : IsAlgebraic K x ↔ IsAlgebraic K (x : L) := (isAlgebraic_algebraMap_iff (algebraMap S L).injective).symm -theorem isIntegral_iff {x : S} : IsIntegral K x ↔ IsIntegral K (x : L) := by - rw [← isAlgebraic_iff_isIntegral, isAlgebraic_iff, isAlgebraic_iff_isIntegral] +theorem isIntegral_iff {x : S} : IsIntegral K x ↔ IsIntegral K (x : L) := + (isIntegral_algHom_iff S.val S.val.injective).symm theorem minpoly_eq (x : S) : minpoly K x = minpoly K (x : L) := (minpoly.algebraMap_eq (algebraMap S L).injective x).symm diff --git a/Mathlib/FieldTheory/IntermediateField/Basic.lean b/Mathlib/FieldTheory/IntermediateField/Basic.lean index 31cc611613755..1c807c5364473 100644 --- a/Mathlib/FieldTheory/IntermediateField/Basic.lean +++ b/Mathlib/FieldTheory/IntermediateField/Basic.lean @@ -6,7 +6,6 @@ Authors: Anne Baanen import Mathlib.Algebra.Field.Subfield import Mathlib.Algebra.Polynomial.AlgebraMap import Mathlib.RingTheory.LocalRing.Basic -import Mathlib.RingTheory.IntegralClosure.IsIntegral.Defs /-! # Intermediate fields @@ -112,10 +111,8 @@ theorem mem_toSubfield (s : IntermediateField K L) (x : L) : x ∈ s.toSubfield definitional equalities. -/ protected def copy (S : IntermediateField K L) (s : Set L) (hs : s = ↑S) : IntermediateField K L where - toSubalgebra := S.toSubalgebra.copy s (hs : s = S.toSubalgebra.carrier) - inv_mem' := - have hs' : (S.toSubalgebra.copy s hs).carrier = S.toSubalgebra.carrier := hs - hs'.symm ▸ S.inv_mem' + toSubalgebra := S.toSubalgebra.copy s hs + inv_mem' := hs.symm ▸ S.inv_mem' @[simp] theorem coe_copy (S : IntermediateField K L) (s : Set L) (hs : s = ↑S) : @@ -341,7 +338,7 @@ instance isScalarTower_mid {R : Type*} [Semiring R] [Algebra L R] [Algebra K R] /-- Specialize `is_scalar_tower_mid` to the common case where the top field is `L` -/ instance isScalarTower_mid' : IsScalarTower K S L := - S.isScalarTower_mid + inferInstance instance {E} [Semiring E] [Algebra L E] : Algebra S E := inferInstanceAs (Algebra S.toSubalgebra E) @@ -468,25 +465,8 @@ instance AlgHom.inhabited : Inhabited (S →ₐ[K] L) := ⟨S.val⟩ theorem aeval_coe {R : Type*} [CommRing R] [Algebra R K] [Algebra R L] [IsScalarTower R K L] - (x : S) (P : R[X]) : aeval (x : L) P = aeval x P := by - refine Polynomial.induction_on' P (fun f g hf hg => ?_) fun n r => ?_ - · rw [aeval_add, aeval_add, AddMemClass.coe_add, hf, hg] - · simp only [MulMemClass.coe_mul, aeval_monomial, SubmonoidClass.coe_pow, mul_eq_mul_right_iff] - left - rfl - -theorem coe_isIntegral_iff {R : Type*} [CommRing R] [Algebra R K] [Algebra R L] - [IsScalarTower R K L] {x : S} : IsIntegral R (x : L) ↔ IsIntegral R x := by - refine ⟨fun h => ?_, fun h => ?_⟩ - · obtain ⟨P, hPmo, hProot⟩ := h - refine ⟨P, hPmo, (injective_iff_map_eq_zero _).1 (algebraMap (↥S) L).injective _ ?_⟩ - letI : IsScalarTower R S L := IsScalarTower.of_algebraMap_eq (congr_fun rfl) - rw [eval₂_eq_eval_map, ← eval₂_at_apply, eval₂_eq_eval_map, Polynomial.map_map, ← - IsScalarTower.algebraMap_eq, ← eval₂_eq_eval_map] - exact hProot - · obtain ⟨P, hPmo, hProot⟩ := h - refine ⟨P, hPmo, ?_⟩ - rw [← aeval_def, aeval_coe, aeval_def, hProot, ZeroMemClass.coe_zero] + (x : S) (P : R[X]) : aeval (x : L) P = aeval x P := + aeval_algHom_apply (S.val.restrictScalars R) x P /-- The map `E → F` when `E` is an intermediate field contained in the intermediate field `F`. diff --git a/Mathlib/FieldTheory/PurelyInseparable.lean b/Mathlib/FieldTheory/PurelyInseparable.lean index 2a1a024969a2d..bf5f64619ceef 100644 --- a/Mathlib/FieldTheory/PurelyInseparable.lean +++ b/Mathlib/FieldTheory/PurelyInseparable.lean @@ -667,7 +667,7 @@ if and only if for any `x ∈ S`, `x ^ (q ^ n)` is contained in `F` for some `n theorem isPurelyInseparable_adjoin_iff_pow_mem (q : ℕ) [hF : ExpChar F q] {S : Set E} : IsPurelyInseparable F (adjoin F S) ↔ ∀ x ∈ S, ∃ n : ℕ, x ^ q ^ n ∈ (algebraMap F E).range := by simp_rw [← le_perfectClosure_iff, adjoin_le_iff, ← mem_perfectClosure_iff_pow_mem q, - Set.le_iff_subset, Set.subset_def, SetLike.mem_coe] + Set.subset_def, SetLike.mem_coe] /-- A compositum of two purely inseparable extensions is purely inseparable. -/ instance isPurelyInseparable_sup (L1 L2 : IntermediateField F E) diff --git a/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean b/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean index 6d30b7aa5aa3a..2bce02c0a50fc 100644 --- a/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean +++ b/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean @@ -671,18 +671,23 @@ noncomputable def divisionRingOfFiniteDimensional (F K : Type*) [Field F] [Ring if H : x = 0 then 0 else Classical.choose <| FiniteDimensional.exists_mul_eq_one F H mul_inv_cancel x hx := show x * dite _ (h := _) _ _ = _ by rw [dif_neg hx] - exact (Classical.choose_spec (FiniteDimensional.exists_mul_eq_one F hx):) + exact (Classical.choose_spec (FiniteDimensional.exists_mul_eq_one F hx) :) inv_zero := dif_pos rfl nnqsmul := _ nnqsmul_def := fun _ _ => rfl qsmul := _ qsmul_def := fun _ _ => rfl +lemma FiniteDimensional.isUnit (F : Type*) {K : Type*} [Field F] [Ring K] [IsDomain K] + [Algebra F K] [FiniteDimensional F K] {x : K} (H : x ≠ 0) : IsUnit x := + let _ := divisionRingOfFiniteDimensional F K; H.isUnit + /-- An integral domain that is module-finite as an algebra over a field is a field. -/ noncomputable def fieldOfFiniteDimensional (F K : Type*) [Field F] [h : CommRing K] [IsDomain K] [Algebra F K] [FiniteDimensional F K] : Field K := { divisionRingOfFiniteDimensional F K with toCommRing := h } + end section DivisionRing diff --git a/Mathlib/RingTheory/Algebraic.lean b/Mathlib/RingTheory/Algebraic.lean index 20fba6c16cb72..8685279579238 100644 --- a/Mathlib/RingTheory/Algebraic.lean +++ b/Mathlib/RingTheory/Algebraic.lean @@ -69,7 +69,7 @@ theorem Algebra.transcendental_iff_not_isAlgebraic : /-- A subalgebra is algebraic if and only if it is algebraic as an algebra. -/ theorem Subalgebra.isAlgebraic_iff (S : Subalgebra R A) : - S.IsAlgebraic ↔ @Algebra.IsAlgebraic R S _ _ S.algebra := by + S.IsAlgebraic ↔ Algebra.IsAlgebraic R S := by delta Subalgebra.IsAlgebraic rw [Subtype.forall', Algebra.isAlgebraic_def] refine forall_congr' fun x => exists_congr fun p => and_congr Iff.rfl ?_ diff --git a/Mathlib/RingTheory/IntegralClosure/Algebra/Basic.lean b/Mathlib/RingTheory/IntegralClosure/Algebra/Basic.lean index 37343ddf2c611..0f9874f3e5ae8 100644 --- a/Mathlib/RingTheory/IntegralClosure/Algebra/Basic.lean +++ b/Mathlib/RingTheory/IntegralClosure/Algebra/Basic.lean @@ -28,15 +28,22 @@ variable {R A B S : Type*} variable [CommRing R] [CommRing A] [Ring B] [CommRing S] variable [Algebra R A] [Algebra R B] (f : R →+* S) +theorem Subalgebra.isIntegral_iff (S : Subalgebra R A) : + Algebra.IsIntegral R S ↔ ∀ x ∈ S, IsIntegral R x := + Algebra.isIntegral_def.trans <| .trans + (forall_congr' fun _ ↦ (isIntegral_algHom_iff S.val Subtype.val_injective).symm) Subtype.forall + section variable {A B : Type*} [Ring A] [Ring B] [Algebra R A] [Algebra R B] -variable (f : A →ₐ[R] B) (hf : Function.Injective f) -theorem Algebra.IsIntegral.of_injective (hf : Function.Injective f) [Algebra.IsIntegral R B] : - Algebra.IsIntegral R A := +theorem Algebra.IsIntegral.of_injective (f : A →ₐ[R] B) (hf : Function.Injective f) + [Algebra.IsIntegral R B] : Algebra.IsIntegral R A := ⟨fun _ ↦ (isIntegral_algHom_iff f hf).mp (isIntegral _)⟩ +theorem AlgEquiv.isIntegral_iff (e : A ≃ₐ[R] B) : Algebra.IsIntegral R A ↔ Algebra.IsIntegral R B := + ⟨fun h ↦ h.of_injective e.symm e.symm.injective, fun h ↦ h.of_injective e e.injective⟩ + end instance Module.End.isIntegral {M : Type*} [AddCommGroup M] [Module R M] [Module.Finite R M] : @@ -57,8 +64,8 @@ instance Algebra.IsIntegral.of_finite [Module.Finite R B] : Algebra.IsIntegral R /-- If `S` is a sub-`R`-algebra of `A` and `S` is finitely-generated as an `R`-module, then all elements of `S` are integral over `R`. -/ -theorem IsIntegral.of_mem_of_fg {A} [Ring A] [Algebra R A] (S : Subalgebra R A) - (HS : S.toSubmodule.FG) (x : A) (hx : x ∈ S) : IsIntegral R x := +theorem IsIntegral.of_mem_of_fg (S : Subalgebra R B) + (HS : S.toSubmodule.FG) (x : B) (hx : x ∈ S) : IsIntegral R x := have : Module.Finite R S := ⟨(fg_top _).mpr HS⟩ (isIntegral_algHom_iff S.val Subtype.val_injective).mpr (.of_finite R (⟨x, hx⟩ : S)) diff --git a/Mathlib/RingTheory/IntegralClosure/IsIntegralClosure/Basic.lean b/Mathlib/RingTheory/IntegralClosure/IsIntegralClosure/Basic.lean index 5c9727459cf1a..59406a95b8ec6 100644 --- a/Mathlib/RingTheory/IntegralClosure/IsIntegralClosure/Basic.lean +++ b/Mathlib/RingTheory/IntegralClosure/IsIntegralClosure/Basic.lean @@ -19,6 +19,58 @@ We prove basic properties of `IsIntegralClosure`. open scoped Classical open Polynomial Submodule +section inv + +open Algebra + +variable {R S : Type*} + +/-- A nonzero element in a domain integral over a field is a unit. -/ +theorem IsIntegral.isUnit [Field R] [Ring S] [IsDomain S] [Algebra R S] {x : S} + (int : IsIntegral R x) (h0 : x ≠ 0) : IsUnit x := + have : FiniteDimensional R (adjoin R {x}) := ⟨(Submodule.fg_top _).mpr int.fg_adjoin_singleton⟩ + (FiniteDimensional.isUnit R (K := adjoin R {x}) + (x := ⟨x, subset_adjoin rfl⟩) <| mt Subtype.ext_iff.mp h0).map (adjoin R {x}).val + +/-- A commutative domain that is an integral algebra over a field is a field. -/ +theorem isField_of_isIntegral_of_isField' [CommRing R] [CommRing S] [IsDomain S] + [Algebra R S] [Algebra.IsIntegral R S] (hR : IsField R) : IsField S where + exists_pair_ne := ⟨0, 1, zero_ne_one⟩ + mul_comm := mul_comm + mul_inv_cancel {x} hx := by + letI := hR.toField + obtain ⟨y, rfl⟩ := (Algebra.IsIntegral.isIntegral (R := R) x).isUnit hx + exact ⟨y.inv, y.val_inv⟩ + +variable [Field R] [DivisionRing S] [Algebra R S] {x : S} {A : Subalgebra R S} + +theorem IsIntegral.inv_mem_adjoin (int : IsIntegral R x) : x⁻¹ ∈ adjoin R {x} := by + obtain rfl | h0 := eq_or_ne x 0 + · rw [inv_zero]; exact Subalgebra.zero_mem _ + have : FiniteDimensional R (adjoin R {x}) := ⟨(Submodule.fg_top _).mpr int.fg_adjoin_singleton⟩ + obtain ⟨⟨y, hy⟩, h1⟩ := FiniteDimensional.exists_mul_eq_one R + (K := adjoin R {x}) (x := ⟨x, subset_adjoin rfl⟩) (mt Subtype.ext_iff.mp h0) + rwa [← mul_left_cancel₀ h0 ((Subtype.ext_iff.mp h1).trans (mul_inv_cancel₀ h0).symm)] + +/-- The inverse of an integral element in a subalgebra of a division ring over a field + also lies in that subalgebra. -/ +theorem IsIntegral.inv_mem (int : IsIntegral R x) (hx : x ∈ A) : x⁻¹ ∈ A := + adjoin_le_iff.mpr (Set.singleton_subset_iff.mpr hx) int.inv_mem_adjoin + +/-- An integral subalgebra of a division ring over a field is closed under inverses. -/ +theorem Algebra.IsIntegral.inv_mem [Algebra.IsIntegral R A] (hx : x ∈ A) : x⁻¹ ∈ A := + ((isIntegral_algHom_iff A.val Subtype.val_injective).mpr <| + Algebra.IsIntegral.isIntegral (⟨x, hx⟩ : A)).inv_mem hx + +/-- The inverse of an integral element in a division ring over a field is also integral. -/ +theorem IsIntegral.inv (int : IsIntegral R x) : IsIntegral R x⁻¹ := + .of_mem_of_fg _ int.fg_adjoin_singleton _ int.inv_mem_adjoin + +theorem IsIntegral.mem_of_inv_mem (int : IsIntegral R x) (inv_mem : x⁻¹ ∈ A) : x ∈ A := by + rw [← inv_inv x]; exact int.inv.inv_mem inv_mem + +end inv + section variable {R A B S : Type*} @@ -78,10 +130,22 @@ theorem le_integralClosure_iff_isIntegral {S : Subalgebra R A} : isIntegral_algebraMap_iff Subtype.coe_injective).trans Algebra.isIntegral_def.symm +theorem Algebra.IsIntegral.adjoin {S : Set A} (hS : ∀ x ∈ S, IsIntegral R x) : + Algebra.IsIntegral R (Algebra.adjoin R S) := + le_integralClosure_iff_isIntegral.mp <| adjoin_le_iff.mpr hS + +theorem integralClosure_eq_top_iff : integralClosure R A = ⊤ ↔ Algebra.IsIntegral R A := by + rw [← top_le_iff, le_integralClosure_iff_isIntegral, + (Subalgebra.topEquiv (R := R) (A := A)).isIntegral_iff] -- explicit arguments for speedup + theorem Algebra.isIntegral_sup {S T : Subalgebra R A} : Algebra.IsIntegral R (S ⊔ T : Subalgebra R A) ↔ Algebra.IsIntegral R S ∧ Algebra.IsIntegral R T := by - simp only [← le_integralClosure_iff_isIntegral, sup_le_iff] + simp_rw [← le_integralClosure_iff_isIntegral, sup_le_iff] + +theorem Algebra.isIntegral_iSup {ι} (S : ι → Subalgebra R A) : + Algebra.IsIntegral R ↑(iSup S) ↔ ∀ i, Algebra.IsIntegral R (S i) := by + simp_rw [← le_integralClosure_iff_isIntegral, iSup_le_iff] /-- Mapping an integral closure along an `AlgEquiv` gives the integral closure. -/ theorem integralClosure_map_algEquiv [Algebra R S] (f : A ≃ₐ[R] S) : @@ -360,6 +424,12 @@ theorem mk'_algebraMap [Algebra R A] [IsScalarTower R A B] (x : R) IsIntegralClosure.mk' A (algebraMap R B x) h = algebraMap R A x := algebraMap_injective A R B <| by rw [algebraMap_mk', ← IsScalarTower.algebraMap_apply] +/-- The integral closure of a field in a commutative domain is always a field. -/ +theorem isField [Algebra R A] [IsScalarTower R A B] [IsDomain A] (hR : IsField R) : + IsField A := + have := IsIntegralClosure.isIntegral_algebra R (A := A) B + isField_of_isIntegral_of_isField' hR + section lift variable (B) {S : Type*} [CommRing S] [Algebra R S] @@ -543,16 +613,6 @@ theorem isField_of_isIntegral_of_isField {R S : Type*} [CommRing R] [CommRing S] ← (injective_iff_map_eq_zero' _).mp hRS, ← aeval_algebraMap_apply_eq_algebraMap_eval] rwa [← eval₂_reverse_eq_zero_iff] at hp -theorem isField_of_isIntegral_of_isField' {R S : Type*} [CommRing R] [CommRing S] [IsDomain S] - [Algebra R S] [Algebra.IsIntegral R S] (hR : IsField R) : IsField S := by - refine ⟨⟨0, 1, zero_ne_one⟩, mul_comm, fun {x} hx ↦ ?_⟩ - have : Module.Finite R (adjoin R {x}) := ⟨(Submodule.fg_top _).mpr - (Algebra.IsIntegral.isIntegral x).fg_adjoin_singleton⟩ - letI := hR.toField - obtain ⟨y, hy⟩ := FiniteDimensional.exists_mul_eq_one R - (K := adjoin R {x}) (x := ⟨x, subset_adjoin rfl⟩) (mt Subtype.ext_iff.mp hx) - exact ⟨y, Subtype.ext_iff.mp hy⟩ - theorem Algebra.IsIntegral.isField_iff_isField {R S : Type*} [CommRing R] [CommRing S] [IsDomain S] [Algebra R S] [Algebra.IsIntegral R S] (hRS : Function.Injective (algebraMap R S)) : IsField R ↔ IsField S := From dc528566839d9a65fe7c84f2c6d2fb17bb49ecbc Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Calle=20S=C3=B6nne?= Date: Thu, 17 Oct 2024 11:36:59 +0000 Subject: [PATCH 072/505] feat(Algebraize): Fix algebraize tactic (#17839) This fixes algebraize tactic to work better when it is pointing to a (possibly custom) constructor/theorem giving the algebra property from the corresponding ringhom property. As an application we apply this to the new `IsStandardSmooth`, `RingHom` property. --- .../RingTheory/RingHom/StandardSmooth.lean | 16 ++++-- Mathlib/Tactic/Algebraize.lean | 50 +++++++------------ test/algebraize.lean | 34 +++++++++++-- 3 files changed, 60 insertions(+), 40 deletions(-) diff --git a/Mathlib/RingTheory/RingHom/StandardSmooth.lean b/Mathlib/RingTheory/RingHom/StandardSmooth.lean index 8b47fbfe7bf3a..4bc826b41f2c7 100644 --- a/Mathlib/RingTheory/RingHom/StandardSmooth.lean +++ b/Mathlib/RingTheory/RingHom/StandardSmooth.lean @@ -31,19 +31,29 @@ namespace RingHom variable {R : Type u} {S : Type v} [CommRing R] [CommRing S] /-- A ring homomorphism `R →+* S` is standard smooth if `S` is standard smooth as `R`-algebra. -/ +@[algebraize RingHom.IsStandardSmooth.toAlgebra] def IsStandardSmooth (f : R →+* S) : Prop := @Algebra.IsStandardSmooth.{t, w} _ _ _ _ f.toAlgebra +/-- Helper lemma for the `algebraize` tactic.-/ +lemma IsStandardSmooth.toAlgebra {f : R →+* S} (hf : IsStandardSmooth.{t, w} f) : + @Algebra.IsStandardSmooth.{t, w} R S _ _ f.toAlgebra := hf + /-- A ring homomorphism `R →+* S` is standard smooth of relative dimension `n` if `S` is standard smooth of relative dimension `n` as `R`-algebra. -/ +@[algebraize RingHom.IsStandardSmoothOfRelativeDimension.toAlgebra] def IsStandardSmoothOfRelativeDimension (f : R →+* S) : Prop := @Algebra.IsStandardSmoothOfRelativeDimension.{t, w} n _ _ _ _ f.toAlgebra +/-- Helper lemma for the `algebraize` tactic.-/ +lemma IsStandardSmoothOfRelativeDimension.toAlgebra {f : R →+* S} + (hf : IsStandardSmoothOfRelativeDimension.{t, w} n f) : + @Algebra.IsStandardSmoothOfRelativeDimension.{t, w} n R S _ _ f.toAlgebra := hf + lemma IsStandardSmoothOfRelativeDimension.isStandardSmooth (f : R →+* S) (hf : IsStandardSmoothOfRelativeDimension.{t, w} n f) : IsStandardSmooth.{t, w} f := by algebraize [f] - letI : Algebra.IsStandardSmoothOfRelativeDimension.{t, w} n R S := hf exact Algebra.IsStandardSmoothOfRelativeDimension.isStandardSmooth n variable {n m} @@ -65,8 +75,6 @@ lemma IsStandardSmooth.comp {g : S →+* T} {f : R →+* S} IsStandardSmooth.{max t t', max w w'} (g.comp f) := by rw [IsStandardSmooth] algebraize [f, g, (g.comp f)] - letI : Algebra.IsStandardSmooth R S := hf - letI : Algebra.IsStandardSmooth S T := hg exact Algebra.IsStandardSmooth.trans.{t, t', w, w'} R S T lemma IsStandardSmoothOfRelativeDimension.comp {g : S →+* T} {f : R →+* S} @@ -75,8 +83,6 @@ lemma IsStandardSmoothOfRelativeDimension.comp {g : S →+* T} {f : R →+* S} IsStandardSmoothOfRelativeDimension.{max t t', max w w'} (n + m) (g.comp f) := by rw [IsStandardSmoothOfRelativeDimension] algebraize [f, g, (g.comp f)] - letI : Algebra.IsStandardSmoothOfRelativeDimension m R S := hf - letI : Algebra.IsStandardSmoothOfRelativeDimension n S T := hg exact Algebra.IsStandardSmoothOfRelativeDimension.trans m n R S T lemma isStandardSmooth_stableUnderComposition : diff --git a/Mathlib/Tactic/Algebraize.lean b/Mathlib/Tactic/Algebraize.lean index 96ceb855e0cb1..1218a802e04ae 100644 --- a/Mathlib/Tactic/Algebraize.lean +++ b/Mathlib/Tactic/Algebraize.lean @@ -38,10 +38,12 @@ specified declaration should be one of the following: 1. An inductive type (i.e. the `Algebra` property itself), in this case it is assumed that the `RingHom` and the `Algebra` property are definitionally the same, and the tactic will construct the -`Algebra` property by giving the `RingHom` property as a term. -2. A constructor for the `Algebra` property. In this case it is assumed that the `RingHom` property -is the last argument of the constructor, and that no other explicit argument is needed. The tactic -then constructs the `Algebra` property by applying the constructor to the `RingHom` property. +`Algebra` property by giving the `RingHom` property as a term. Due to how this is peformed, we also +need to assume that the `Algebra` property can be constructed only from the homomorphism, so it can +not have any other explicit arguments. +2. A lemma (or constructor) proving the `Algebra` property from the `RingHom` property. In this case +it is assumed that the `RingHom` property is the final argument, and that no other explicit argument +is needed. The tactic then constructs the `Algebra` property by applying the lemma or constructor. Here are three examples of properties tagged with the `algebraize` attribute: ``` @@ -98,9 +100,9 @@ There are two cases for what declaration corresponding to this `Name` can be. 1. An inductive type (i.e. the `Algebra` property itself), in this case it is assumed that the `RingHom` and the `Algebra` property are definitionally the same, and the tactic will construct the `Algebra` property by giving the `RingHom` property as a term. -2. A constructor for the `Algebra` property. In this case it is assumed that the `RingHom` property -is the last argument of the constructor, and that no other explicit argument is needed. The tactic -then constructs the `Algebra` property by applying the constructor to the `RingHom` property. +2. A lemma (or constructor) proving the `Algebra` property from the `RingHom` property. In this case +it is assumed that the `RingHom` property is the final argument, and that no other explicit argument +is needed. The tactic then constructs the `Algebra` property by applying the lemma or constructor. Finally, if no argument is provided to the `algebraize` attribute, it is assumed that the tagged declaration has name `RingHom.Property` and that the corresponding `Algebra` property has name @@ -171,7 +173,8 @@ def addProperties (t : Array Expr) : TacticM Unit := withMainContext do let (nm, args) := decl.type.getAppFnArgs -- Check if the type of the current hypothesis has been tagged with the `algebraize` attribute match Attr.algebraizeAttr.getParam? (← getEnv) nm with - -- If it has, `p` will be the name of the corresponding `Algebra` property (or a constructor) + -- If it has, `p` will either be the name of the corresponding `Algebra` property, or a + -- lemma/constructor. | some p => -- The last argument of the `RingHom` property is assumed to be `f` let f := args[args.size - 1]! @@ -184,42 +187,27 @@ def addProperties (t : Array Expr) : TacticM Unit := withMainContext do /- If the attribute points to the corresponding `Algebra` property itself, we assume that it is definitionally the same as the `RingHom` property. Then, we just need to construct its type and the local declaration will already give a valid term. -/ - match cinfo with - | .inductInfo _ => + if cinfo.isInductive then let pargs := pargs.set! 0 args[0]! let pargs := pargs.set! 1 args[1]! let tp ← mkAppOptM p pargs -- This should be the type `Algebra.Property A B` unless (← synthInstance? tp).isSome do liftMetaTactic fun mvarid => do let nm ← mkFreshBinderNameForTactic `algebraizeInst - let mvar ← mvarid.define nm tp decl.toExpr - let (_, mvar) ← mvar.intro1P + let (_, mvar) ← mvarid.note nm decl.toExpr tp return [mvar] - /- Otherwise, the attribute points to a constructor of the `Algebra` property. In this case, - we assume that the `RingHom` property is the last argument of the constructor (and that - this is all we need to supply explicitly). -/ - | .ctorInfo ctor => - -- construct the desired value + /- Otherwise, the attribute points to a lemma or a constructor for the `Algebra` property. + In this case, we assume that the `RingHom` property is the last argument of the lemma or + constructor (and that this is all we need to supply explicitly). -/ + else let pargs := pargs.set! (n - 1) decl.toExpr let val ← mkAppOptM p pargs - - -- construct the expected type - let alg ← mkAppOptM ``Algebra #[args[0]!, args[1]!, none, none] - let algInst := (← synthInstance? alg) - let mut argsType := Array.mkArray (ctor.numParams) (none : Option Expr) - argsType := argsType.set! 0 args[0]! - argsType := argsType.set! 1 args[1]! - argsType := argsType.set! (ctor.numParams - 1) algInst - let tp := ← mkAppOptM ctor.induct argsType - + let tp ← inferType val unless (← synthInstance? tp).isSome do liftMetaTactic fun mvarid => do let nm ← mkFreshBinderNameForTactic `algebraizeInst - let mvar ← mvarid.define nm tp val - let (_, mvar) ← mvar.intro1P + let (_, mvar) ← mvarid.note nm val return [mvar] - | _ => logError s!"bad argument to `algebraize` attribute: {p}. \ - Only supporting inductive types or constructors." | none => return /-- Configuration for `algebraize`. -/ diff --git a/test/algebraize.lean b/test/algebraize.lean index c2aa17892ab57..1ccd695d0b920 100644 --- a/test/algebraize.lean +++ b/test/algebraize.lean @@ -38,6 +38,26 @@ an example. -/ def RingHom.testProperty3 {A B : Type*} [CommRing A] [CommRing B] (f : A →+* B) : Prop := f.testProperty1 +/- Test property for when the `RingHom` (and `Algebra`) property have an extra explicit argument, +and hence needs to be created through a lemma. See e.g. +`Algebra.IsStandardSmoothOfRelativeDimension` for an example. -/ +class Algebra.testProperty4 (n : ℕ) (A B : Type*) [CommRing A] [CommRing B] [Algebra A B] : Prop where + out : ∀ m, n = m + +/- Test property for when the `RingHom` (and `Algebra`) property have an extra explicit argument, +and hence needs to be created through a lemma. See e.g. +`Algebra.IsStandardSmoothOfRelativeDimension` for an example. -/ +@[algebraize testProperty4.toAlgebra] +def RingHom.testProperty4 (n : ℕ) {A B : Type*} [CommRing A] [CommRing B] (_ : A →+* B) : Prop := + ∀ m, n = m + +lemma testProperty4.toAlgebra (n : ℕ) {A B : Type*} [CommRing A] [CommRing B] (f : A →+* B) + (hf : f.testProperty4 n) : + letI : Algebra A B := f.toAlgebra + Algebra.testProperty4 n A B := + letI : Algebra A B := f.toAlgebra + { out := hf } + end example_definitions set_option tactic.hygienic false @@ -79,17 +99,23 @@ example (A B C : Type*) [CommRing A] [CommRing B] [CommRing C] (f : A →+* B) ( example (A B : Type*) [CommRing A] [CommRing B] (f : A →+* B) (hf : f.testProperty1) : True := by algebraize [f] - guard_hyp algebraizeInst : Algebra.testProperty1 A B := hf + guard_hyp algebraizeInst : Algebra.testProperty1 A B trivial example (A B : Type*) [CommRing A] [CommRing B] (f : A →+* B) (hf : f.testProperty2) : True := by algebraize [f] - guard_hyp algebraizeInst : Module.testProperty2 A B := hf + guard_hyp algebraizeInst : Module.testProperty2 A B trivial example (A B : Type*) [CommRing A] [CommRing B] (f : A →+* B) (hf : f.testProperty3) : True := by algebraize [f] - guard_hyp algebraizeInst : Algebra.testProperty3 A B := ⟨hf⟩ + guard_hyp algebraizeInst : Algebra.testProperty3 A B + trivial + +example (n : ℕ) (A B : Type*) [CommRing A] [CommRing B] (f : A →+* B) (hf : f.testProperty4 n) : + True := by + algebraize [f] + guard_hyp algebraizeInst : Algebra.testProperty4 n A B trivial /-- Synthesize from morphism property of a composition (and check that tower is also synthesized). -/ @@ -100,6 +126,6 @@ example (A B C : Type*) [CommRing A] [CommRing B] [CommRing C] (f : A →+* B) ( fail_if_success have h : IsScalarTower A B C := inferInstance algebraize [f, g, g.comp f] - guard_hyp algebraizeInst : Algebra.testProperty1 A C := hfg + guard_hyp algebraizeInst : Algebra.testProperty1 A C guard_hyp scalarTowerInst := IsScalarTower.of_algebraMap_eq' rfl trivial From 9266f86d046ce02c551f1d9fe9f425dabb44476f Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Thu, 17 Oct 2024 12:44:26 +0000 Subject: [PATCH 073/505] feat(RingTheory/Flat/FaithfullyFlat): definition of faithfully flat module (#17805) 1. $M$ is $R$-faithfully flat iff $M$ is flat and $IM \ne M$ for all maximal ideal $I \le R$ 2. prove that this is equivalent to that $M$ is flat and $M \otimes N \ne 0$ whenever $N \ne 0$ --- Mathlib.lean | 1 + Mathlib/RingTheory/Flat/FaithfullyFlat.lean | 168 ++++++++++++++++++++ 2 files changed, 169 insertions(+) create mode 100644 Mathlib/RingTheory/Flat/FaithfullyFlat.lean diff --git a/Mathlib.lean b/Mathlib.lean index 0439d6b8d5e5f..8c55dc3701c47 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3941,6 +3941,7 @@ import Mathlib.RingTheory.Flat.Algebra import Mathlib.RingTheory.Flat.Basic import Mathlib.RingTheory.Flat.CategoryTheory import Mathlib.RingTheory.Flat.EquationalCriterion +import Mathlib.RingTheory.Flat.FaithfullyFlat import Mathlib.RingTheory.Flat.Stability import Mathlib.RingTheory.FractionalIdeal.Basic import Mathlib.RingTheory.FractionalIdeal.Extended diff --git a/Mathlib/RingTheory/Flat/FaithfullyFlat.lean b/Mathlib/RingTheory/Flat/FaithfullyFlat.lean new file mode 100644 index 0000000000000..17c82bb7c375d --- /dev/null +++ b/Mathlib/RingTheory/Flat/FaithfullyFlat.lean @@ -0,0 +1,168 @@ +/- +Copyright (c) 2024 Judith Ludwig, Florent Schaffhauser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Judith Ludwig, Florent Schaffhauser, Yunzhou Xie, Jujian Zhang +-/ + +import Mathlib.RingTheory.Flat.Basic +import Mathlib.LinearAlgebra.TensorProduct.Quotient + +/-! +# Faithfully flat modules + +A module `M` over a commutative ring `R` is *faithfully flat* if it is flat and `IM ≠ M` whenever +`I` is a maximal ideal of `R`. + +## Main declaration + +- `Module.FaithfullyFlat`: the predicate asserting that an `R`-module `M` is faithfully flat. + +## Main theorems + +- `Module.FaithfullyFlat.iff_flat_and_proper_ideal`: an `R`-module `M` is faithfully flat iff it is + flat and for all proper ideals `I` of `R`, `I • M ≠ M`. +- `Module.FaithfullyFlat.iff_flat_and_rTensor_faithful`: an `R`-module `M` is faithfully flat iff it + is flat and tensoring with `M` is faithful, i.e. `N ≠ 0` implies `N ⊗ M ≠ 0`. +- `Module.FaithfullyFlat.iff_flat_and_lTensor_faithful`: an `R`-module `M` is faithfully flat iff it + is flat and tensoring with `M` is faithful, i.e. `N ≠ 0` implies `M ⊗ N ≠ 0`. + +- `Module.FaithfullyFlat.self`: the `R`-module `R` is faithfully flat. + +-/ + +universe u v + +open TensorProduct + +namespace Module + +variable (R : Type u) (M : Type v) [CommRing R] [AddCommGroup M] [Module R M] + +/-- +A module `M` over a commutative ring `R` is *faithfully flat* if it is flat and, +for all `R`-module homomorphism `f : N → N'` such that `id ⊗ f = 0`, we have `f = 0`. +-/ +@[mk_iff] class FaithfullyFlat extends Module.Flat R M : Prop where + submodule_ne_top : ∀ ⦃m : Ideal R⦄ (_ : Ideal.IsMaximal m), m • (⊤ : Submodule R M) ≠ ⊤ + +namespace FaithfullyFlat +instance self : FaithfullyFlat R R where + submodule_ne_top m h r := Ideal.eq_top_iff_one _ |>.not.1 h.ne_top <| by + simpa using show 1 ∈ (m • ⊤ : Ideal R) from r.symm ▸ ⟨⟩ + +section proper_ideal + +lemma iff_flat_and_proper_ideal : + FaithfullyFlat R M ↔ + (Flat R M ∧ ∀ (I : Ideal R), I ≠ ⊤ → I • (⊤ : Submodule R M) ≠ ⊤) := by + rw [faithfullyFlat_iff] + refine ⟨fun ⟨flat, h⟩ => ⟨flat, fun I hI r => ?_⟩, fun h => ⟨h.1, fun m hm => h.2 _ hm.ne_top⟩⟩ + obtain ⟨m, hm, le⟩ := I.exists_le_maximal hI + exact h hm <| eq_top_iff.2 <| show ⊤ ≤ m • ⊤ from r ▸ Submodule.smul_mono le (by simp [r]) + +lemma iff_flat_and_ideal_smul_eq_top : + FaithfullyFlat R M ↔ + (Flat R M ∧ ∀ (I : Ideal R), I • (⊤ : Submodule R M) = ⊤ → I = ⊤) := + iff_flat_and_proper_ideal R M |>.trans <| and_congr_right_iff.2 fun _ => iff_of_eq <| + forall_congr fun I => eq_iff_iff.2 <| by tauto + +end proper_ideal + +section faithful + +instance rTensor_nontrivial + [fl: FaithfullyFlat R M] (N : Type*) [AddCommGroup N] [Module R N] [Nontrivial N] : + Nontrivial (N ⊗[R] M) := by + obtain ⟨n, hn⟩ := nontrivial_iff_exists_ne (0 : N) |>.1 inferInstance + let I := (Submodule.span R {n}).annihilator + by_cases I_ne_top : I = ⊤ + · rw [Ideal.eq_top_iff_one, Submodule.mem_annihilator_span_singleton, one_smul] at I_ne_top + contradiction + let inc : R ⧸ I →ₗ[R] N := Submodule.liftQ _ ((LinearMap.lsmul R N).flip n) <| fun r hr => by + simpa only [LinearMap.mem_ker, LinearMap.flip_apply, LinearMap.lsmul_apply, + Submodule.mem_annihilator_span_singleton, I] using hr + have injective_inc : Function.Injective inc := LinearMap.ker_eq_bot.1 <| eq_bot_iff.2 <| by + intro r hr + induction r using Quotient.inductionOn' with | h r => + simpa only [Submodule.Quotient.mk''_eq_mk, Submodule.mem_bot, Submodule.Quotient.mk_eq_zero, + Submodule.mem_annihilator_span_singleton, LinearMap.mem_ker, Submodule.liftQ_apply, + LinearMap.flip_apply, LinearMap.lsmul_apply, I, inc] using hr + have ne_top := iff_flat_and_proper_ideal R M |>.1 fl |>.2 I I_ne_top + refine subsingleton_or_nontrivial _ |>.resolve_left fun rid => ?_ + exact False.elim <| ne_top <| Submodule.subsingleton_quotient_iff_eq_top.1 <| + Function.Injective.comp (g := LinearMap.rTensor M inc) + (fl.toFlat.rTensor_preserves_injective_linearMap inc injective_inc) + ((quotTensorEquivQuotSMul M I).symm.injective) |>.subsingleton + +instance lTensor_nontrivial + [FaithfullyFlat R M] (N : Type*) [AddCommGroup N] [Module R N] [Nontrivial N] : + Nontrivial (M ⊗[R] N) := + TensorProduct.comm R M N |>.toEquiv.nontrivial + +lemma rTensor_reflects_triviality + [FaithfullyFlat R M] (N : Type*) [AddCommGroup N] [Module R N] + [h : Subsingleton (N ⊗[R] M)] : Subsingleton N := by + revert h; change _ → _; contrapose + simp only [not_subsingleton_iff_nontrivial] + intro h + infer_instance + +lemma lTensor_reflects_triviality + [FaithfullyFlat R M] (N : Type*) [AddCommGroup N] [Module R N] + [Subsingleton (M ⊗[R] N)] : + Subsingleton N := by + haveI : Subsingleton (N ⊗[R] M) := (TensorProduct.comm R N M).toEquiv.injective.subsingleton + apply rTensor_reflects_triviality R M + +attribute [-simp] Ideal.Quotient.mk_eq_mk in +lemma iff_flat_and_rTensor_faithful : + FaithfullyFlat R M ↔ + (Flat R M ∧ + ∀ (N : Type max u v) [AddCommGroup N] [Module R N], + Nontrivial N → Nontrivial (N ⊗[R] M)) := by + refine ⟨fun fl => ⟨inferInstance, rTensor_nontrivial R M⟩, fun ⟨flat, faithful⟩ => ⟨?_⟩⟩ + intro m hm rid + specialize faithful (ULift (R ⧸ m)) inferInstance + haveI : Nontrivial ((R ⧸ m) ⊗[R] M) := + (congr (ULift.moduleEquiv : ULift (R ⧸ m) ≃ₗ[R] R ⧸ m) + (LinearEquiv.refl R M)).symm.toEquiv.nontrivial + have := (quotTensorEquivQuotSMul M m).toEquiv.symm.nontrivial + haveI H : Subsingleton (M ⧸ m • (⊤ : Submodule R M)) := by + rwa [Submodule.subsingleton_quotient_iff_eq_top] + rw [← not_nontrivial_iff_subsingleton] at H + contradiction + +lemma iff_flat_and_rTensor_reflects_triviality : + FaithfullyFlat R M ↔ + (Flat R M ∧ + ∀ (N : Type max u v) [AddCommGroup N] [Module R N], + Subsingleton (N ⊗[R] M) → Subsingleton N) := + iff_flat_and_rTensor_faithful R M |>.trans <| and_congr_right_iff.2 fun _ => iff_of_eq <| + forall_congr fun N => forall_congr fun _ => forall_congr fun _ => iff_iff_eq.1 <| by + simp only [← not_subsingleton_iff_nontrivial]; tauto + +lemma iff_flat_and_lTensor_faithful : + FaithfullyFlat R M ↔ + (Flat R M ∧ + ∀ (N : Type max u v) [AddCommGroup N] [Module R N], + Nontrivial N → Nontrivial (M ⊗[R] N)) := + iff_flat_and_rTensor_faithful R M |>.trans + ⟨fun ⟨flat, faithful⟩ => ⟨flat, fun N _ _ _ => + letI := faithful N inferInstance; (TensorProduct.comm R M N).toEquiv.nontrivial⟩, + fun ⟨flat, faithful⟩ => ⟨flat, fun N _ _ _ => + letI := faithful N inferInstance; (TensorProduct.comm R M N).symm.toEquiv.nontrivial⟩⟩ + +lemma iff_flat_and_lTensor_reflects_triviality : + FaithfullyFlat R M ↔ + (Flat R M ∧ + ∀ (N : Type max u v) [AddCommGroup N] [Module R N], + Subsingleton (M ⊗[R] N) → Subsingleton N) := + iff_flat_and_lTensor_faithful R M |>.trans <| and_congr_right_iff.2 fun _ => iff_of_eq <| + forall_congr fun N => forall_congr fun _ => forall_congr fun _ => iff_iff_eq.1 <| by + simp only [← not_subsingleton_iff_nontrivial]; tauto + +end faithful + +end FaithfullyFlat + +end Module From fa6b2491c26ecf81fcfe648152d07d9faff7a9f0 Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Thu, 17 Oct 2024 12:44:27 +0000 Subject: [PATCH 074/505] chore(LinearAlgebra/LinearIndependent): generalize a section from rings to semirings (#17864) Co-authored-by: Eric Wieser --- Mathlib/LinearAlgebra/LinearIndependent.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/LinearAlgebra/LinearIndependent.lean b/Mathlib/LinearAlgebra/LinearIndependent.lean index c198ef6a8918d..ccd018f0d3ddf 100644 --- a/Mathlib/LinearAlgebra/LinearIndependent.lean +++ b/Mathlib/LinearAlgebra/LinearIndependent.lean @@ -1183,8 +1183,8 @@ end Module section Nontrivial -variable [Ring R] [Nontrivial R] [AddCommGroup M] [AddCommGroup M'] -variable [Module R M] [NoZeroSMulDivisors R M] [Module R M'] +variable [Semiring R] [Nontrivial R] [AddCommMonoid M] +variable [Module R M] [NoZeroSMulDivisors R M] variable {s t : Set M} theorem linearIndependent_unique_iff (v : ι → M) [Unique ι] : From 0c022ea19db4295764843578cfbcb3aae9347bbb Mon Sep 17 00:00:00 2001 From: Scott Carnahan <128885296+ScottCarnahan@users.noreply.github.com> Date: Thu, 17 Oct 2024 13:07:51 +0000 Subject: [PATCH 075/505] feat (RingTheory/HahnSeries) : coefficient-wise map (#16164) This PR introduces `HahnSeries.map` which applies a zero-preserving map to each coefficient. We more or less copy the `Polynomial.map` API. --- Mathlib/RingTheory/HahnSeries/Addition.lean | 39 ++++++++++++++----- Mathlib/RingTheory/HahnSeries/Basic.lean | 29 +++++++++++--- .../RingTheory/HahnSeries/Multiplication.lean | 26 ++++++++++++- .../RingTheory/HahnSeries/PowerSeries.lean | 6 +-- 4 files changed, 81 insertions(+), 19 deletions(-) diff --git a/Mathlib/RingTheory/HahnSeries/Addition.lean b/Mathlib/RingTheory/HahnSeries/Addition.lean index 3daeb21892dce..255b47ebace23 100644 --- a/Mathlib/RingTheory/HahnSeries/Addition.lean +++ b/Mathlib/RingTheory/HahnSeries/Addition.lean @@ -28,7 +28,7 @@ open scoped Classical noncomputable section -variable {Γ R : Type*} +variable {Γ Γ' R S U V : Type*} namespace HahnSeries @@ -66,6 +66,10 @@ theorem add_coeff' {x y : HahnSeries Γ R} : (x + y).coeff = x.coeff + y.coeff : theorem add_coeff {x y : HahnSeries Γ R} {a : Γ} : (x + y).coeff a = x.coeff a + y.coeff a := rfl +@[simp] +protected lemma map_add [AddMonoid S] (f : R →+ S) {x y : HahnSeries Γ R} : + ((x + y).map f : HahnSeries Γ S) = x.map f + y.map f := by + ext; simp /-- `addOppositeEquiv` is an additive monoid isomorphism between Hahn series over `Γ` with coefficients in the opposite additive monoid `Rᵃᵒᵖ` @@ -210,7 +214,7 @@ def coeff.addMonoidHom (g : Γ) : HahnSeries Γ R →+ R where section Domain -variable {Γ' : Type*} [PartialOrder Γ'] +variable [PartialOrder Γ'] theorem embDomain_add (f : Γ ↪o Γ') (x y : HahnSeries Γ R) : embDomain f (x + y) = embDomain f x + embDomain f y := by @@ -261,12 +265,9 @@ theorem support_neg {x : HahnSeries Γ R} : (-x).support = x.support := by simp @[simp] -theorem sub_coeff' {x y : HahnSeries Γ R} : (x - y).coeff = x.coeff - y.coeff := by - ext - simp [sub_eq_add_neg] - -theorem sub_coeff {x y : HahnSeries Γ R} {a : Γ} : (x - y).coeff a = x.coeff a - y.coeff a := by - simp +protected lemma map_neg [AddGroup S] (f : R →+ S) {x : HahnSeries Γ R} : + ((-x).map f : HahnSeries Γ S) = -(x.map f) := by + ext; simp theorem orderTop_neg {x : HahnSeries Γ R} : (-x).orderTop = x.orderTop := by simp only [orderTop, support_neg, neg_eq_zero] @@ -277,6 +278,19 @@ theorem order_neg [Zero Γ] {f : HahnSeries Γ R} : (-f).order = f.order := by · simp only [hf, neg_zero] simp only [order, support_neg, neg_eq_zero] +@[simp] +theorem sub_coeff' {x y : HahnSeries Γ R} : (x - y).coeff = x.coeff - y.coeff := by + ext + simp [sub_eq_add_neg] + +theorem sub_coeff {x y : HahnSeries Γ R} {a : Γ} : (x - y).coeff a = x.coeff a - y.coeff a := by + simp + +@[simp] +protected lemma map_sub [AddGroup S] (f : R →+ S) {x y : HahnSeries Γ R} : + ((x - y).map f : HahnSeries Γ S) = x.map f - y.map f := by + ext; simp + theorem min_orderTop_le_orderTop_sub {Γ} [LinearOrder Γ] {x y : HahnSeries Γ R} : min x.orderTop y.orderTop ≤ (x - y).orderTop := by rw [sub_eq_add_neg, ← orderTop_neg (x := y)] @@ -377,7 +391,7 @@ end DistribMulAction section Module -variable [PartialOrder Γ] [Semiring R] {V : Type*} [AddCommMonoid V] [Module R V] +variable [PartialOrder Γ] [Semiring R] [AddCommMonoid V] [Module R V] instance : Module R (HahnSeries Γ V) := { inferInstanceAs (DistribMulAction R (HahnSeries Γ V)) with @@ -401,9 +415,14 @@ def single.linearMap (a : Γ) : V →ₗ[R] HahnSeries Γ V := def coeff.linearMap (g : Γ) : HahnSeries Γ V →ₗ[R] V := { coeff.addMonoidHom g with map_smul' := fun _ _ => rfl } +@[simp] +protected lemma map_smul [AddCommMonoid U] [Module R U] (f : U →ₗ[R] V) {r : R} + {x : HahnSeries Γ U} : (r • x).map f = r • ((x.map f) : HahnSeries Γ V) := by + ext; simp + section Domain -variable {Γ' : Type*} [PartialOrder Γ'] +variable [PartialOrder Γ'] theorem embDomain_smul (f : Γ ↪o Γ') (r : R) (x : HahnSeries Γ R) : embDomain f (r • x) = r • embDomain f x := by diff --git a/Mathlib/RingTheory/HahnSeries/Basic.lean b/Mathlib/RingTheory/HahnSeries/Basic.lean index 29586b2758fad..7c4261e42f42c 100644 --- a/Mathlib/RingTheory/HahnSeries/Basic.lean +++ b/Mathlib/RingTheory/HahnSeries/Basic.lean @@ -25,6 +25,8 @@ in the file `RingTheory/LaurentSeries`. coefficient if `x ≠ 0`, and is `⊤` when `x = 0`. * `order x` is a minimal element of `Γ` where `x` has a nonzero coefficient if `x ≠ 0`, and is zero when `x = 0`. +* `map` takes each coefficient of a Hahn series to its target under a zero-preserving map. +* `embDomain` preserves coefficients, but embeds the index set `Γ` in a larger poset. ## References - [J. van der Hoeven, *Operators on Generalized Power Series*][van_der_hoeven] @@ -43,7 +45,7 @@ structure HahnSeries (Γ : Type*) (R : Type*) [PartialOrder Γ] [Zero R] where coeff : Γ → R isPWO_support' : (Function.support coeff).IsPWO -variable {Γ : Type*} {R : Type*} +variable {Γ Γ' R S : Type*} namespace HahnSeries @@ -108,8 +110,20 @@ nonrec theorem support_nonempty_iff {x : HahnSeries Γ R} : x.support.Nonempty theorem support_eq_empty_iff {x : HahnSeries Γ R} : x.support = ∅ ↔ x = 0 := Function.support_eq_empty_iff.trans coeff_fun_eq_zero_iff +/-- The map of Hahn series induced by applying a zero-preserving map to each coefficient. -/ +@[simps] +def map [Zero S] (x : HahnSeries Γ R) {F : Type*} [FunLike F R S] [ZeroHomClass F R S] (f : F) : + HahnSeries Γ S where + coeff g := f (x.coeff g) + isPWO_support' := x.isPWO_support.mono <| Function.support_comp_subset (ZeroHomClass.map_zero f) _ + +@[simp] +protected lemma map_zero [Zero S] (f : ZeroHom R S) : + (0 : HahnSeries Γ R).map f = 0 := by + ext; simp + /-- Change a HahnSeries with coefficients in HahnSeries to a HahnSeries on the Lex product. -/ -def ofIterate {Γ' : Type*} [PartialOrder Γ'] (x : HahnSeries Γ (HahnSeries Γ' R)) : +def ofIterate [PartialOrder Γ'] (x : HahnSeries Γ (HahnSeries Γ' R)) : HahnSeries (Γ ×ₗ Γ') R where coeff := fun g => coeff (coeff x g.1) g.2 isPWO_support' := by @@ -125,7 +139,7 @@ lemma mk_eq_zero (f : Γ → R) (h) : HahnSeries.mk f h = 0 ↔ f = 0 := by rfl /-- Change a Hahn series on a lex product to a Hahn series with coefficients in a Hahn series. -/ -def toIterate {Γ' : Type*} [PartialOrder Γ'] (x : HahnSeries (Γ ×ₗ Γ') R) : +def toIterate [PartialOrder Γ'] (x : HahnSeries (Γ ×ₗ Γ') R) : HahnSeries Γ (HahnSeries Γ' R) where coeff := fun g => { coeff := fun g' => coeff x (g, g') @@ -141,7 +155,7 @@ def toIterate {Γ' : Type*} [PartialOrder Γ'] (x : HahnSeries (Γ ×ₗ Γ') R) /-- The equivalence between iterated Hahn series and Hahn series on the lex product. -/ @[simps] -def iterateEquiv {Γ' : Type*} [PartialOrder Γ'] : +def iterateEquiv [PartialOrder Γ'] : HahnSeries Γ (HahnSeries Γ' R) ≃ HahnSeries (Γ ×ₗ Γ') R where toFun := ofIterate invFun := toIterate @@ -194,6 +208,11 @@ theorem single_ne_zero (h : r ≠ 0) : single a r ≠ 0 := fun con => theorem single_eq_zero_iff {a : Γ} {r : R} : single a r = 0 ↔ r = 0 := map_eq_zero_iff _ <| single_injective a +@[simp] +protected lemma map_single [Zero S] (f : ZeroHom R S) : (single a r).map f = single a (f r) := by + ext g + by_cases h : g = a <;> simp [h] + instance [Nonempty Γ] [Nontrivial R] : Nontrivial (HahnSeries Γ R) := ⟨by obtain ⟨r, s, rs⟩ := exists_pair_ne R @@ -364,7 +383,7 @@ end Order section Domain -variable {Γ' : Type*} [PartialOrder Γ'] +variable [PartialOrder Γ'] open Classical in /-- Extends the domain of a `HahnSeries` by an `OrderEmbedding`. -/ diff --git a/Mathlib/RingTheory/HahnSeries/Multiplication.lean b/Mathlib/RingTheory/HahnSeries/Multiplication.lean index 11ac4db05c9d9..8f62c2deef2dc 100644 --- a/Mathlib/RingTheory/HahnSeries/Multiplication.lean +++ b/Mathlib/RingTheory/HahnSeries/Multiplication.lean @@ -37,7 +37,7 @@ open Finset Function Pointwise noncomputable section -variable {Γ Γ' R V : Type*} +variable {Γ Γ' R S V : Type*} namespace HahnSeries @@ -74,6 +74,12 @@ theorem order_one [MulZeroOneClass R] : order (1 : HahnSeries Γ R) = 0 := by theorem leadingCoeff_one [MulZeroOneClass R] : (1 : HahnSeries Γ R).leadingCoeff = 1 := by simp [leadingCoeff_eq] +@[simp] +protected lemma map_one [MonoidWithZero R] [MonoidWithZero S] (f : R →*₀ S) : + (1 : HahnSeries Γ R).map f = (1 : HahnSeries Γ S) := by + ext g + by_cases h : g = 0 <;> simp [h] + end HahnSeries /-- We introduce a type alias for `HahnSeries` in order to work with scalar multiplication by @@ -348,6 +354,19 @@ theorem mul_coeff [NonUnitalNonAssocSemiring R] {x y : HahnSeries Γ R} {a : Γ} ∑ ij ∈ addAntidiagonal x.isPWO_support y.isPWO_support a, x.coeff ij.fst * y.coeff ij.snd := rfl +protected lemma map_mul [NonUnitalNonAssocSemiring R] [NonUnitalNonAssocSemiring S] (f : R →ₙ+* S) + {x y : HahnSeries Γ R} : (x * y).map f = (x.map f : HahnSeries Γ S) * (y.map f) := by + ext + simp only [map_coeff, mul_coeff, ZeroHom.coe_coe, map_sum, map_mul] + refine Eq.symm (sum_subset (fun gh hgh => ?_) (fun gh hgh hz => ?_)) + · simp_all only [mem_addAntidiagonal, mem_support, map_coeff, ZeroHom.coe_coe, ne_eq, and_true] + exact ⟨fun h => hgh.1 (map_zero f ▸ congrArg f h), fun h => hgh.2.1 (map_zero f ▸ congrArg f h)⟩ + · simp_all only [mem_addAntidiagonal, mem_support, ne_eq, map_coeff, ZeroHom.coe_coe, and_true, + not_and, not_not] + by_cases h : f (x.coeff gh.1) = 0 + · exact mul_eq_zero_of_left h (f (y.coeff gh.2)) + · exact mul_eq_zero_of_right (f (x.coeff gh.1)) (hz h) + theorem mul_coeff_left' [NonUnitalNonAssocSemiring R] {x y : HahnSeries Γ R} {a : Γ} {s : Set Γ} (hs : s.IsPWO) (hxs : x.support ⊆ s) : (x * y).coeff a = @@ -642,6 +661,11 @@ theorem C_zero : C (0 : R) = (0 : HahnSeries Γ R) := theorem C_one : C (1 : R) = (1 : HahnSeries Γ R) := C.map_one +theorem map_C [NonAssocSemiring S] (a : R) (f : R →+* S) : + ((C a).map f : HahnSeries Γ S) = C (f a) := by + ext g + by_cases h : g = 0 <;> simp [h] + theorem C_injective : Function.Injective (C : R → HahnSeries Γ R) := by intro r s rs rw [HahnSeries.ext_iff, funext_iff] at rs diff --git a/Mathlib/RingTheory/HahnSeries/PowerSeries.lean b/Mathlib/RingTheory/HahnSeries/PowerSeries.lean index 094a71acddd64..7a70d9e119bf8 100644 --- a/Mathlib/RingTheory/HahnSeries/PowerSeries.lean +++ b/Mathlib/RingTheory/HahnSeries/PowerSeries.lean @@ -37,7 +37,7 @@ open Finset Function Pointwise Polynomial noncomputable section -variable {Γ : Type*} {R : Type*} +variable {Γ R : Type*} namespace HahnSeries @@ -113,7 +113,7 @@ theorem ofPowerSeries_C (r : R) : ofPowerSeries Γ R (PowerSeries.C R r) = HahnS single_coeff] split_ifs with hn · subst hn - convert @embDomain_coeff ℕ R _ _ Γ _ _ _ 0 <;> simp + convert embDomain_coeff (a := 0) <;> simp · rw [embDomain_notin_image_support] simp only [not_exists, Set.mem_image, toPowerSeries_symm_apply_coeff, mem_support, PowerSeries.coeff_C] @@ -126,7 +126,7 @@ theorem ofPowerSeries_X : ofPowerSeries Γ R PowerSeries.X = single 1 1 := by simp only [single_coeff, ofPowerSeries_apply, RingHom.coe_mk] split_ifs with hn · rw [hn] - convert @embDomain_coeff ℕ R _ _ Γ _ _ _ 1 <;> simp + convert embDomain_coeff (a := 1) <;> simp · rw [embDomain_notin_image_support] simp only [not_exists, Set.mem_image, toPowerSeries_symm_apply_coeff, mem_support, PowerSeries.coeff_X] From a0d3126672d97e23df807041f3960f375244c9a6 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Thu, 17 Oct 2024 13:07:53 +0000 Subject: [PATCH 076/505] feat(RingTheory/Kaehler): the first homology of the cotangent complex (#17559) Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- .../RingTheory/Kaehler/CotangentComplex.lean | 132 ++++++++++++++++-- 1 file changed, 122 insertions(+), 10 deletions(-) diff --git a/Mathlib/RingTheory/Kaehler/CotangentComplex.lean b/Mathlib/RingTheory/Kaehler/CotangentComplex.lean index 75e250102f8fe..cd899098a4f46 100644 --- a/Mathlib/RingTheory/Kaehler/CotangentComplex.lean +++ b/Mathlib/RingTheory/Kaehler/CotangentComplex.lean @@ -22,6 +22,10 @@ defined by `I`), we may define the (naive) cotangent complex `I/I² → ⨁ᵢ S - `Algebra.Generators.exact_cotangentComplex_toKaehler`: `I/I² → ⨁ᵢ S dxᵢ → Ω[S/R]` is exact. - `Algebra.Generators.Hom.Sub`: If `f` and `g` are two maps between presentations, `f - g` induces a map `⨁ᵢ S dxᵢ → I/I²` that makes `f` and `g` homotopic. +- `Algebra.Generators.H1Cotangent`: The first homology of the (naive) cotangent complex + of `S` over `R`, induced by a given presentation. +- `Algebra.H1Cotangent`: `H¹(L_{S/R})`, + the first homology of the (naive) cotangent complex of `S` over `R`. -/ open KaehlerDifferential TensorProduct MvPolynomial @@ -83,6 +87,16 @@ attribute [local instance] SMulCommClass.of_commMonoid variable {P P'} +universe w'' u'' v'' + +variable {R'' : Type u''} {S'' : Type v''} [CommRing R''] [CommRing S''] [Algebra R'' S''] +variable {P'' : Generators.{w''} R'' S''} +variable [Algebra R R''] [Algebra S S''] [Algebra R S''] + [IsScalarTower R R'' S''] [IsScalarTower R S S''] +variable [Algebra R' R''] [Algebra S' S''] [Algebra R' S''] + [IsScalarTower R' R'' S''] [IsScalarTower R' S' S''] +variable [IsScalarTower S S' S''] + namespace CotangentSpace /-- @@ -114,16 +128,6 @@ lemma repr_map (f : Hom P P') (i j) : simp only [cotangentSpaceBasis_apply, map_tmul, map_one, Hom.toAlgHom_X, cotangentSpaceBasis_repr_one_tmul] -universe w'' u'' v'' - -variable {R'' : Type u''} {S'' : Type v''} [CommRing R''] [CommRing S''] [Algebra R'' S''] -variable {P'' : Generators.{w''} R'' S''} -variable [Algebra R R''] [Algebra S S''] [Algebra R S''] - [IsScalarTower R R'' S''] [IsScalarTower R S S''] -variable [Algebra R' R''] [Algebra S' S''] [Algebra R' S''] - [IsScalarTower R' R'' S''] [IsScalarTower R' S' S''] -variable [IsScalarTower S S' S''] - @[simp] lemma map_id : CotangentSpace.map (.id P) = LinearMap.id := by @@ -275,6 +279,114 @@ lemma toKaehler_surjective : Function.Surjective P.toKaehler := lemma exact_cotangentComplex_toKaehler : Function.Exact P.cotangentComplex P.toKaehler := exact_kerCotangentToTensor_mapBaseChange _ _ _ P.algebraMap_surjective +variable (P) in +/-- +The first homology of the (naive) cotangent complex of `S` over `R`, +induced by a given presentation `0 → I → P → R → 0`, +defined as the kernel of `I/I² → S ⊗[P] Ω[P⁄R]`. +-/ +protected noncomputable +def H1Cotangent : Type _ := LinearMap.ker P.cotangentComplex + +variable {P : Generators R S} + +noncomputable +instance : AddCommGroup P.H1Cotangent := by delta Generators.H1Cotangent; infer_instance + +noncomputable +instance {R₀} [CommRing R₀] [Algebra R₀ S] [Module R₀ P.Cotangent] + [IsScalarTower R₀ S P.Cotangent] : Module R₀ P.H1Cotangent := by + delta Generators.H1Cotangent; infer_instance + +@[simp] lemma H1Cotangent.val_add (x y : P.H1Cotangent) : (x + y).1 = x.1 + y.1 := rfl +@[simp] lemma H1Cotangent.val_zero : (0 : P.H1Cotangent).1 = 0 := rfl +@[simp] lemma H1Cotangent.val_smul {R₀} [CommRing R₀] [Algebra R₀ S] [Module R₀ P.Cotangent] + [IsScalarTower R₀ S P.Cotangent] (r : R₀) (x : P.H1Cotangent) : (r • x).1 = r • x.1 := rfl + +noncomputable +instance {R₁ R₂} [CommRing R₁] [CommRing R₂] [Algebra R₁ R₂] + [Algebra R₁ S] [Algebra R₂ S] + [Module R₁ P.Cotangent] [IsScalarTower R₁ S P.Cotangent] + [Module R₂ P.Cotangent] [IsScalarTower R₂ S P.Cotangent] + [IsScalarTower R₁ R₂ P.Cotangent] : + IsScalarTower R₁ R₂ P.H1Cotangent := by + delta Generators.H1Cotangent; infer_instance + +lemma subsingleton_h1Cotangent (P : Generators R S) : + Subsingleton P.H1Cotangent ↔ Function.Injective P.cotangentComplex := by + delta Generators.H1Cotangent + rw [← LinearMap.ker_eq_bot, Submodule.eq_bot_iff, subsingleton_iff_forall_eq 0, Subtype.forall'] + simp only [Subtype.ext_iff, Submodule.coe_zero] + +/-- The inclusion of `H¹(L_{S/R})` into the conormal space of a presentation. -/ +@[simps!] def h1Cotangentι : P.H1Cotangent →ₗ[S] P.Cotangent := Submodule.subtype _ + +lemma h1Cotangentι_injective : Function.Injective P.h1Cotangentι := Subtype.val_injective + +@[ext] lemma h1Cotangentι_ext (x y : P.H1Cotangent) (e : x.1 = y.1) : x = y := Subtype.ext e + +/-- +The induced map on the first homology of the (naive) cotangent complex. +-/ +@[simps!] +noncomputable +def H1Cotangent.map (f : Hom P P') : P.H1Cotangent →ₗ[S] P'.H1Cotangent := by + refine (Cotangent.map f).restrict (p := LinearMap.ker P.cotangentComplex) + (q := (LinearMap.ker P'.cotangentComplex).restrictScalars S) fun x hx ↦ ?_ + simp only [LinearMap.mem_ker, Submodule.restrictScalars_mem] at hx ⊢ + apply_fun (CotangentSpace.map f) at hx + rw [CotangentSpace.map_cotangentComplex] at hx + rw [hx] + exact LinearMap.map_zero _ + +lemma H1Cotangent.map_eq (f g : Hom P P') : map f = map g := by + ext x + simp only [map_apply_coe] + rw [← sub_eq_zero, ← Cotangent.val_sub, ← LinearMap.sub_apply, Cotangent.map_sub_map] + simp only [LinearMap.coe_comp, Function.comp_apply, LinearMap.map_coe_ker, map_zero, + Cotangent.val_zero] + +@[simp] lemma H1Cotangent.map_id : map (.id P) = LinearMap.id := by ext; simp + +lemma H1Cotangent.map_comp [IsScalarTower R R' R''] (f : Hom P P') (g : Hom P' P'') : + map (g.comp f) = (map g).restrictScalars S ∘ₗ map f := by ext; simp [Cotangent.map_comp] + +/-- `H¹(L_{S/R})` is independent of the presentation chosen. -/ +@[simps! apply] +noncomputable +def H1Cotangent.equiv (P : Generators.{w} R S) (P' : Generators.{w'} R S) : + P.H1Cotangent ≃ₗ[S] P'.H1Cotangent where + __ := map default + invFun := map default + left_inv x := + show ((map (defaultHom P' P)) ∘ₗ (map (defaultHom P P'))) x = LinearMap.id x by + rw [← map_id, eq_comm, map_eq _ ((defaultHom P' P).comp (defaultHom P P')), map_comp]; rfl + right_inv x := + show ((map (defaultHom P P')) ∘ₗ (map (defaultHom P' P))) x = LinearMap.id x by + rw [← map_id, eq_comm, map_eq _ ((defaultHom P P').comp (defaultHom P' P)), map_comp]; rfl + end Generators +variable {S' : Type*} [CommRing S'] [Algebra R S'] +variable {T : Type w} [CommRing T] [Algebra R T] [Algebra S T] [IsScalarTower R S T] +variable [Algebra S' T] [IsScalarTower R S' T] + +variable (R S S' T) + +/-- `H¹(L_{S/R})`, the first homology of the (naive) cotangent complex of `S` over `R`. -/ +abbrev H1Cotangent : Type _ := (Generators.self R S).H1Cotangent + +/-- The induced map on the first homology of the (naive) cotangent complex of `S` over `R`. -/ +noncomputable +def H1Cotangent.map : H1Cotangent R S' →ₗ[S'] H1Cotangent S T := + Generators.H1Cotangent.map (Generators.defaultHom _ _) + +variable {R S S' T} + +/-- `H¹(L_{S/R})` is independent of the presentation chosen. -/ +noncomputable +abbrev Generators.equivH1Cotangent (P : Generators.{w} R S) : + P.H1Cotangent ≃ₗ[S] H1Cotangent R S := + Generators.H1Cotangent.equiv _ _ + end Algebra From 83062057427a0227a6dbb7b07622f1f7e64b7dc2 Mon Sep 17 00:00:00 2001 From: mathlib4-update-dependencies-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Thu, 17 Oct 2024 13:07:54 +0000 Subject: [PATCH 077/505] chore: update Mathlib dependencies 2024-10-17 (#17865) This PR updates the Mathlib dependencies. --- lake-manifest.json | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lake-manifest.json b/lake-manifest.json index 1dc73cf0afe47..b86cb2dfc481e 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,7 +5,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "cc0bc876eeef0518ddc1c8d3bd6f48cc83e68901", + "rev": "c521f0185f4dd42b6aa4898010d5ba5357c57c9f", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", From 090e2437af8d263f97b731cf78d012f45e0be700 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 17 Oct 2024 13:44:43 +0000 Subject: [PATCH 078/505] =?UTF-8?q?feat:=20`PosSMulMono=20=E2=84=9A=20?= =?UTF-8?q?=CE=B1=20=E2=86=92=20PosSMulMono=20=E2=84=9A=E2=89=A50=20=CE=B1?= =?UTF-8?q?`=20(#17217)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit From LeanAPAP --- Mathlib/Algebra/Order/Module/Rat.lean | 6 ++++++ 1 file changed, 6 insertions(+) diff --git a/Mathlib/Algebra/Order/Module/Rat.lean b/Mathlib/Algebra/Order/Module/Rat.lean index 0368ae68efa89..288226c006963 100644 --- a/Mathlib/Algebra/Order/Module/Rat.lean +++ b/Mathlib/Algebra/Order/Module/Rat.lean @@ -12,6 +12,12 @@ import Mathlib.Data.Rat.Cast.Order variable {α : Type*} +instance PosSMulMono.nnrat_of_rat [Preorder α] [MulAction ℚ α] [PosSMulMono ℚ α] : + PosSMulMono ℚ≥0 α where elim _q hq _a₁ _a₂ ha := smul_le_smul_of_nonneg_left (α := ℚ) ha hq + +instance PosSMulStrictMono.nnrat_of_rat [Preorder α] [MulAction ℚ α] [PosSMulStrictMono ℚ α] : + PosSMulStrictMono ℚ≥0 α where elim _q hq _a₁ _a₂ ha := smul_lt_smul_of_pos_left (α := ℚ) ha hq + section LinearOrderedAddCommGroup variable [LinearOrderedAddCommGroup α] From 4dfa80d4a2b80ed9668f60f14cc147cb335a1c23 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Thu, 17 Oct 2024 14:00:37 +0000 Subject: [PATCH 079/505] feat: Polynomial.degree_prod_of_monic (#17330) --- Mathlib/Algebra/Polynomial/BigOperators.lean | 11 +++++++++++ 1 file changed, 11 insertions(+) diff --git a/Mathlib/Algebra/Polynomial/BigOperators.lean b/Mathlib/Algebra/Polynomial/BigOperators.lean index cef12c507490a..909aed4291d6a 100644 --- a/Mathlib/Algebra/Polynomial/BigOperators.lean +++ b/Mathlib/Algebra/Polynomial/BigOperators.lean @@ -193,10 +193,21 @@ theorem natDegree_multiset_prod_of_monic (h : ∀ f ∈ t, Monic f) : assumption · simp +theorem degree_multiset_prod_of_monic [Nontrivial R] (h : ∀ f ∈ t, Monic f) : + t.prod.degree = (t.map degree).sum := by + have : t.prod ≠ 0 := Monic.ne_zero <| by simpa using monic_multiset_prod_of_monic _ _ h + rw [degree_eq_natDegree this, natDegree_multiset_prod_of_monic _ h, Nat.cast_multiset_sum, + Multiset.map_map, Function.comp_def, + Multiset.map_congr rfl (fun f hf => (degree_eq_natDegree (h f hf).ne_zero).symm)] + theorem natDegree_prod_of_monic (h : ∀ i ∈ s, (f i).Monic) : (∏ i ∈ s, f i).natDegree = ∑ i ∈ s, (f i).natDegree := by simpa using natDegree_multiset_prod_of_monic (s.1.map f) (by simpa using h) +theorem degree_prod_of_monic [Nontrivial R] (h : ∀ i ∈ s, (f i).Monic) : + (∏ i ∈ s, f i).degree = ∑ i ∈ s, (f i).degree := by + simpa using degree_multiset_prod_of_monic (s.1.map f) (by simpa using h) + theorem coeff_multiset_prod_of_natDegree_le (n : ℕ) (hl : ∀ p ∈ t, natDegree p ≤ n) : coeff t.prod ((Multiset.card t) * n) = (t.map fun p => coeff p n).prod := by induction t using Quotient.inductionOn From d82cdbff79fdde07da9d4edc86a3d9d4fbc67c13 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Thu, 17 Oct 2024 14:30:25 +0000 Subject: [PATCH 080/505] feat(RingTheory/Smooth): Smoothness of product rings. (#15131) --- Mathlib.lean | 1 + Mathlib/Algebra/Algebra/Pi.lean | 9 ++ .../RingTheory/Ideal/QuotientOperations.lean | 6 + Mathlib/RingTheory/Smooth/Pi.lean | 121 ++++++++++++++++++ 4 files changed, 137 insertions(+) create mode 100644 Mathlib/RingTheory/Smooth/Pi.lean diff --git a/Mathlib.lean b/Mathlib.lean index 8c55dc3701c47..155e528def04f 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4133,6 +4133,7 @@ import Mathlib.RingTheory.SimpleRing.Basic import Mathlib.RingTheory.SimpleRing.Defs import Mathlib.RingTheory.Smooth.Basic import Mathlib.RingTheory.Smooth.Kaehler +import Mathlib.RingTheory.Smooth.Pi import Mathlib.RingTheory.Smooth.StandardSmooth import Mathlib.RingTheory.Support import Mathlib.RingTheory.SurjectiveOnStalks diff --git a/Mathlib/Algebra/Algebra/Pi.lean b/Mathlib/Algebra/Algebra/Pi.lean index db4ea723211c1..fbdfdb6ca061d 100644 --- a/Mathlib/Algebra/Algebra/Pi.lean +++ b/Mathlib/Algebra/Algebra/Pi.lean @@ -52,6 +52,15 @@ theorem algebraMap_apply {_ : CommSemiring R} [_s : ∀ i, Semiring (f i)] [∀ -- when each `A i` is an `R i`-algebra, although I'm not sure that it's useful. variable {I} (R) +/-- A family of algebra homomorphisms `g i : A →ₐ[R] f i` defines a ring homomorphism +`Pi.algHom g : A →ₐ[R] Π i, f i` given by `Pi.algHom g x i = f i x`. -/ +@[simps!] +def algHom [CommSemiring R] [s : ∀ i, Semiring (f i)] [∀ i, Algebra R (f i)] + {A : Type*} [Semiring A] [Algebra R A] (g : ∀ i, A →ₐ[R] f i) : + A →ₐ[R] ∀ i, f i where + __ := Pi.ringHom fun i ↦ (g i).toRingHom + commutes' r := by ext; simp + /-- `Function.eval` as an `AlgHom`. The name matches `Pi.evalRingHom`, `Pi.evalMonoidHom`, etc. -/ @[simps] diff --git a/Mathlib/RingTheory/Ideal/QuotientOperations.lean b/Mathlib/RingTheory/Ideal/QuotientOperations.lean index 5fb624f6f2143..efc97f56b8993 100644 --- a/Mathlib/RingTheory/Ideal/QuotientOperations.lean +++ b/Mathlib/RingTheory/Ideal/QuotientOperations.lean @@ -479,6 +479,12 @@ theorem quotientMap_comp_mk {J : Ideal R} {I : Ideal S} {f : R →+* S} (H : J (quotientMap I f H).comp (Quotient.mk J) = (Quotient.mk I).comp f := RingHom.ext fun x => by simp only [Function.comp_apply, RingHom.coe_comp, Ideal.quotientMap_mk] +lemma ker_quotientMap_mk {I J : Ideal R} : + RingHom.ker (quotientMap (J.map _) (Quotient.mk I) le_comap_map) = I.map (Quotient.mk J) := by + rw [Ideal.quotientMap, Ideal.ker_quotient_lift, ← RingHom.comap_ker, Ideal.mk_ker, + Ideal.comap_map_of_surjective _ Ideal.Quotient.mk_surjective, + ← RingHom.ker_eq_comap_bot, Ideal.mk_ker, Ideal.map_sup, Ideal.map_quotient_self, bot_sup_eq] + /-- The ring equiv `R/I ≃+* S/J` induced by a ring equiv `f : R ≃+* S`, where `J = f(I)`. -/ @[simps] def quotientEquiv (I : Ideal R) (J : Ideal S) (f : R ≃+* S) (hIJ : J = I.map (f : R →+* S)) : diff --git a/Mathlib/RingTheory/Smooth/Pi.lean b/Mathlib/RingTheory/Smooth/Pi.lean new file mode 100644 index 0000000000000..acebf93edbce4 --- /dev/null +++ b/Mathlib/RingTheory/Smooth/Pi.lean @@ -0,0 +1,121 @@ +/- +Copyright (c) 2024 Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Andrew Yang +-/ +import Mathlib.RingTheory.Idempotents +import Mathlib.RingTheory.Smooth.Basic + +/-! + +# Formal-smoothness of finite products of rings + +## Main result + +- `Algebra.FormallySmooth.pi_iff`: If `I` is finite, `Π i : I, A i` is `R`-formally-smooth + if and only if each `A i` is `R`-formally-smooth. + +-/ + +universe u v + +namespace Algebra.FormallySmooth + +variable {R : Type (max u v)} {I : Type u} (A : I → Type (max u v)) +variable [CommRing R] [∀ i, CommRing (A i)] [∀ i, Algebra R (A i)] + +theorem of_pi [FormallySmooth R (Π i, A i)] (i) : + FormallySmooth R (A i) := by + classical + fapply FormallySmooth.of_split (Pi.evalAlgHom R A i) + apply AlgHom.ofLinearMap + ((Ideal.Quotient.mkₐ R _).toLinearMap.comp (LinearMap.single _ _ i)) + · show Ideal.Quotient.mk _ (Pi.single i 1) = 1 + rw [← (Ideal.Quotient.mk _).map_one, ← sub_eq_zero, ← map_sub, + Ideal.Quotient.eq_zero_iff_mem] + have : Pi.single i 1 - 1 ∈ RingHom.ker (Pi.evalAlgHom R A i).toRingHom := by + simp [RingHom.mem_ker] + convert neg_mem (Ideal.pow_mem_pow this 2) using 1 + simp [pow_two, sub_mul, mul_sub, ← Pi.single_mul] + · intro x y + show Ideal.Quotient.mk _ _ = Ideal.Quotient.mk _ _ * Ideal.Quotient.mk _ _ + simp only [AlgHom.toRingHom_eq_coe, LinearMap.coe_single, Pi.single_mul, map_mul] + · ext x + show (Pi.single i x) i = x + simp + +theorem pi_iff [Finite I] : + FormallySmooth R (Π i, A i) ↔ ∀ i, FormallySmooth R (A i) := by + classical + cases nonempty_fintype I + constructor + · exact fun _ ↦ of_pi A + · intro H + constructor + intros B _ _ J hJ g + have hJ' (x) (hx : x ∈ RingHom.ker (Ideal.Quotient.mk J)) : IsNilpotent x := by + refine ⟨2, show x ^ 2 ∈ (⊥ : Ideal B) from ?_⟩ + rw [← hJ] + exact Ideal.pow_mem_pow (by simpa using hx) 2 + obtain ⟨e, he, he'⟩ := ((CompleteOrthogonalIdempotents.single A).map + g.toRingHom).lift_of_isNilpotent_ker (Ideal.Quotient.mk J) hJ' + fun _ ↦ Ideal.Quotient.mk_surjective _ + replace he' : ∀ i, Ideal.Quotient.mk J (e i) = g (Pi.single i 1) := congr_fun he' + let iso : B ≃ₐ[R] ∀ i, B ⧸ Ideal.span {1 - e i} := + { __ := Pi.algHom _ _ fun i ↦ Ideal.Quotient.mkₐ R _ + __ := Equiv.ofBijective _ he.bijective_pi } + let J' := fun i ↦ J.map (Ideal.Quotient.mk (Ideal.span {1 - e i})) + let ι : ∀ i, (B ⧸ J →ₐ[R] (B ⧸ _) ⧸ J' i) := fun i ↦ Ideal.quotientMapₐ _ + (IsScalarTower.toAlgHom R B _) Ideal.le_comap_map + have hι : ∀ i x, ι i x = 0 → (e i) * x = 0 := by + intros i x hix + have : x ∈ (Ideal.span {1 - e i}).map (Ideal.Quotient.mk J) := by + rw [← Ideal.ker_quotientMap_mk]; exact hix + rw [Ideal.map_span, Set.image_singleton, Ideal.mem_span_singleton] at this + obtain ⟨c, rfl⟩ := this + rw [← mul_assoc, ← map_mul, mul_sub, mul_one, (he.idem i).eq, sub_self, map_zero, zero_mul] + have : ∀ i : I, ∃ a : A i →ₐ[R] B ⧸ Ideal.span {1 - e i}, ∀ x, + Ideal.Quotient.mk (J' i) (a x) = ι i (g (Pi.single i x)) := by + intro i + let g' : A i →ₐ[R] (B ⧸ _) ⧸ (J' i) := by + apply AlgHom.ofLinearMap (((ι i).comp g).toLinearMap ∘ₗ LinearMap.single _ _ i) + · suffices Ideal.Quotient.mk (Ideal.span {1 - e i}) (e i) = 1 by simp [ι, ← he', this] + rw [← (Ideal.Quotient.mk _).map_one, eq_comm, Ideal.Quotient.mk_eq_mk_iff_sub_mem, + Ideal.mem_span_singleton] + · intros x y; simp [Pi.single_mul] + obtain ⟨a, ha⟩ := FormallySmooth.comp_surjective (I := J' i) + (by rw [← Ideal.map_pow, hJ, Ideal.map_bot]) g' + exact ⟨a, AlgHom.congr_fun ha⟩ + choose a ha using this + use iso.symm.toAlgHom.comp (Pi.algHom _ _ fun i ↦ (a i).comp (Pi.evalAlgHom R A i)) + ext x; rw [← AlgHom.toLinearMap_apply, ← AlgHom.toLinearMap_apply]; congr 1 + ext i x + simp only [AlgEquiv.toAlgHom_eq_coe, AlgHom.comp_toLinearMap, AlgEquiv.toAlgHom_toLinearMap, + LinearMap.coe_comp, LinearMap.coe_single, Function.comp_apply, AlgHom.toLinearMap_apply, + AlgEquiv.toLinearMap_apply, Ideal.Quotient.mkₐ_eq_mk] + obtain ⟨y, hy⟩ := Ideal.Quotient.mk_surjective (a i x) + have hy' : Ideal.Quotient.mk (Ideal.span {1 - e i}) (y * e i) = a i x := by + have : Ideal.Quotient.mk (Ideal.span {1 - e i}) (e i) = 1 := by + rw [← (Ideal.Quotient.mk _).map_one, eq_comm, Ideal.Quotient.mk_eq_mk_iff_sub_mem, + Ideal.mem_span_singleton] + rw [map_mul, this, hy, mul_one] + trans Ideal.Quotient.mk J (y * e i) + · congr 1; apply iso.injective; ext j + suffices a j (Pi.single i x j) = Ideal.Quotient.mk _ (y * e i) by simpa using this + by_cases hij : i = j + · subst hij + rw [Pi.single_eq_same, hy'] + · have : Ideal.Quotient.mk (Ideal.span {1 - e j}) (e i) = 0 := by + rw [Ideal.Quotient.eq_zero_iff_mem, Ideal.mem_span_singleton] + refine ⟨e i, by simp [he.ortho (Ne.symm hij), sub_mul]⟩ + rw [Pi.single_eq_of_ne (Ne.symm hij), map_zero, map_mul, this, mul_zero] + · have : ι i (Ideal.Quotient.mk J (y * e i)) = ι i (g (Pi.single i x)) := by + rw [← ha, ← hy'] + simp only [Ideal.quotient_map_mkₐ, IsScalarTower.coe_toAlgHom', + Ideal.Quotient.algebraMap_eq, Ideal.Quotient.mkₐ_eq_mk, ι] + rw [← sub_eq_zero, ← map_sub] at this + replace this := hι _ _ this + rwa [mul_sub, ← map_mul, mul_comm, mul_assoc, (he.idem i).eq, he', ← map_mul, ← Pi.single_mul, + one_mul, sub_eq_zero] at this + +end Algebra.FormallySmooth From e2b89093a1a30606a1de57c265dd8d8d33d3ae2b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 17 Oct 2024 14:30:27 +0000 Subject: [PATCH 081/505] chore: don't export `GroupWithZero.inv_zero` (#17465) Exporting makes it not show as `inv_zero` in the docs search results, meaning it's very non-discoverable. See https://github.com/leanprover/doc-gen4/issues/222. --- Mathlib/Algebra/GroupWithZero/Defs.lean | 7 +++---- 1 file changed, 3 insertions(+), 4 deletions(-) diff --git a/Mathlib/Algebra/GroupWithZero/Defs.lean b/Mathlib/Algebra/GroupWithZero/Defs.lean index 14ace092e132d..221f060ebda9a 100644 --- a/Mathlib/Algebra/GroupWithZero/Defs.lean +++ b/Mathlib/Algebra/GroupWithZero/Defs.lean @@ -182,16 +182,15 @@ Examples include division rings and the ordered monoids that are the target of valuations in general valuation theory. -/ class GroupWithZero (G₀ : Type u) extends MonoidWithZero G₀, DivInvMonoid G₀, Nontrivial G₀ where /-- The inverse of `0` in a group with zero is `0`. -/ - inv_zero : (0 : G₀)⁻¹ = 0 + protected inv_zero : (0 : G₀)⁻¹ = 0 /-- Every nonzero element of a group with zero is invertible. -/ protected mul_inv_cancel (a : G₀) : a ≠ 0 → a * a⁻¹ = 1 -export GroupWithZero (inv_zero) -attribute [simp] inv_zero - section GroupWithZero variable [GroupWithZero G₀] {a : G₀} +@[simp] lemma inv_zero : (0 : G₀)⁻¹ = 0 := GroupWithZero.inv_zero + @[simp] lemma mul_inv_cancel₀ (h : a ≠ 0) : a * a⁻¹ = 1 := GroupWithZero.mul_inv_cancel a h -- See note [lower instance priority] From 58282561d2e928698a6727596a4ad96d74f64766 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 17 Oct 2024 14:30:28 +0000 Subject: [PATCH 082/505] feat: Lagrange's theorem for the pointwise product of a subgroup with a set (#17680) From LeanCamCombi --- Mathlib/GroupTheory/Coset/Card.lean | 11 +++++++++++ 1 file changed, 11 insertions(+) diff --git a/Mathlib/GroupTheory/Coset/Card.lean b/Mathlib/GroupTheory/Coset/Card.lean index 67a686a06e9b1..df9e2e1b9a667 100644 --- a/Mathlib/GroupTheory/Coset/Card.lean +++ b/Mathlib/GroupTheory/Coset/Card.lean @@ -10,6 +10,8 @@ import Mathlib.SetTheory.Cardinal.Finite # Lagrange's theorem: the order of a subgroup divides the order of the group. -/ +open scoped Pointwise + namespace Subgroup variable {α : Type*} [Group α] @@ -19,6 +21,15 @@ theorem card_eq_card_quotient_mul_card_subgroup (s : Subgroup α) : Nat.card α = Nat.card (α ⧸ s) * Nat.card s := by rw [← Nat.card_prod]; exact Nat.card_congr Subgroup.groupEquivQuotientProdSubgroup +@[to_additive] +lemma card_mul_eq_card_subgroup_mul_card_quotient (s : Subgroup α) (t : Set α) : + Nat.card (t * s : Set α) = Nat.card s * Nat.card (t.image (↑) : Set (α ⧸ s)) := by + rw [← Nat.card_prod, Nat.card_congr] + apply Equiv.trans _ (QuotientGroup.preimageMkEquivSubgroupProdSet _ _) + rw [QuotientGroup.preimage_image_mk] + convert Equiv.refl ↑(t * s) + aesop (add simp [Set.mem_mul]) + /-- **Lagrange's Theorem**: The order of a subgroup divides the order of its ambient group. -/ @[to_additive "**Lagrange's Theorem**: The order of an additive subgroup divides the order of its ambient additive group."] From 436ba3afd94da896fa1d20b511a1469ea5d018c3 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Thu, 17 Oct 2024 14:30:30 +0000 Subject: [PATCH 083/505] chore: remove "instance was not necessary" porting notes (issue #10670) (#17866) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Discussed with @kim-em. These instances are now necessary because instance synthesis now searches with the full discrimination tree, so mismatches in the arguments to instances are scrutinized much harder. This is not likely to be reverted in the future, so the instances remain required. There were a few further occurrences of #10670 for Grp, see e.g. [Grp.instCoeFunHomForallαGroup](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Algebra/Category/Grp/Basic.html#Grp.instCoeFunHomForall%CE%B1Group). These do look like they are unnecessary but removing them breaks a lot of downstream proofs. (I think it's something to do with `coe_of` unfolding `of` causing a mismatch between the expected types, but it's a bit too subtle to deal with easily.) I'll leave those porting notes for the meantime. We need a big refactor cleaning up the concrete category bits of the library anyway. --- Mathlib/Algebra/Category/Grp/Basic.lean | 2 -- Mathlib/Algebra/Category/Grp/Preadditive.lean | 1 - Mathlib/Geometry/RingedSpace/Basic.lean | 1 - Mathlib/Order/Category/NonemptyFinLinOrd.lean | 1 - 4 files changed, 5 deletions(-) diff --git a/Mathlib/Algebra/Category/Grp/Basic.lean b/Mathlib/Algebra/Category/Grp/Basic.lean index 67d5908ffff48..1c03372df70b3 100644 --- a/Mathlib/Algebra/Category/Grp/Basic.lean +++ b/Mathlib/Algebra/Category/Grp/Basic.lean @@ -112,7 +112,6 @@ instance hasForgetToMonCat : HasForget₂ Grp MonCat := @[to_additive] instance : Coe Grp.{u} MonCat.{u} where coe := (forget₂ Grp MonCat).obj --- porting note (#10670): this instance was not necessary in mathlib @[to_additive] instance (G H : Grp) : One (G ⟶ H) := (inferInstance : One (MonoidHom G H)) @@ -257,7 +256,6 @@ instance hasForgetToCommMonCat : HasForget₂ CommGrp CommMonCat := @[to_additive] instance : Coe CommGrp.{u} CommMonCat.{u} where coe := (forget₂ CommGrp CommMonCat).obj --- porting note (#10670): this instance was not necessary in mathlib @[to_additive] instance (G H : CommGrp) : One (G ⟶ H) := (inferInstance : One (MonoidHom G H)) diff --git a/Mathlib/Algebra/Category/Grp/Preadditive.lean b/Mathlib/Algebra/Category/Grp/Preadditive.lean index 5866733e35b85..fd7ca0d120ef3 100644 --- a/Mathlib/Algebra/Category/Grp/Preadditive.lean +++ b/Mathlib/Algebra/Category/Grp/Preadditive.lean @@ -16,7 +16,6 @@ universe u namespace AddCommGrp --- porting note (#10670): this instance was not necessary in mathlib instance (P Q : AddCommGrp) : AddCommGroup (P ⟶ Q) := (inferInstance : AddCommGroup (AddMonoidHom P Q)) diff --git a/Mathlib/Geometry/RingedSpace/Basic.lean b/Mathlib/Geometry/RingedSpace/Basic.lean index 4a89dde97d9c9..a11af973bc890 100644 --- a/Mathlib/Geometry/RingedSpace/Basic.lean +++ b/Mathlib/Geometry/RingedSpace/Basic.lean @@ -48,7 +48,6 @@ lemma res_zero {X : RingedSpace.{u}} {U V : TopologicalSpace.Opens X} variable (X : RingedSpace) --- Porting note (#10670): this was not necessary in mathlib3 instance : CoeSort RingedSpace Type* where coe X := X.carrier diff --git a/Mathlib/Order/Category/NonemptyFinLinOrd.lean b/Mathlib/Order/Category/NonemptyFinLinOrd.lean index cfa1dab425200..29ef2f1dc61d2 100644 --- a/Mathlib/Order/Category/NonemptyFinLinOrd.lean +++ b/Mathlib/Order/Category/NonemptyFinLinOrd.lean @@ -121,7 +121,6 @@ instance {A B : NonemptyFinLinOrd.{u}} : FunLike (A ⟶ B) A B where ext x exact congr_fun h x --- porting note (#10670): this instance was not necessary in mathlib instance {A B : NonemptyFinLinOrd.{u}} : OrderHomClass (A ⟶ B) A B where map_rel f _ _ h := f.monotone h From 12a15447c89cc5bbbfa031af07a954ae31b3bbdf Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Thu, 17 Oct 2024 14:30:31 +0000 Subject: [PATCH 084/505] chore: remove "CoeFun -> FunLike" porting/adaptation notes (#17871) We prefer `FunLike` over `CoeFun` if possible, so we don't care that we made that change during the port. See also #17870 for applying the change that was porting noted consistently across the entirety of Mathlib. --- Mathlib/Algebra/Module/LinearMap/Defs.lean | 2 -- Mathlib/Algebra/MvPolynomial/Basic.lean | 2 -- Mathlib/Analysis/InnerProductSpace/PiL2.lean | 1 - Mathlib/LinearAlgebra/Multilinear/Basic.lean | 1 - Mathlib/NumberTheory/ArithmeticFunction.lean | 1 - Mathlib/Order/Heyting/Hom.lean | 7 ------- Mathlib/Order/Hom/CompleteLattice.lean | 6 ------ Mathlib/RingTheory/RingInvo.lean | 4 ---- Mathlib/Topology/Algebra/Module/Basic.lean | 9 --------- Mathlib/Topology/Category/TopCat/Basic.lean | 6 ------ 10 files changed, 39 deletions(-) diff --git a/Mathlib/Algebra/Module/LinearMap/Defs.lean b/Mathlib/Algebra/Module/LinearMap/Defs.lean index eceefd6b766ab..0a75f33f5bf9c 100644 --- a/Mathlib/Algebra/Module/LinearMap/Defs.lean +++ b/Mathlib/Algebra/Module/LinearMap/Defs.lean @@ -216,8 +216,6 @@ lemma coe_coe {F : Type*} [FunLike F M M₃] [SemilinearMapClass F σ M M₃] {f ⇑(f : M →ₛₗ[σ] M₃) = f := rfl --- Porting note: we don't port specialized `CoeFun` instances if there is `DFunLike` instead - /-- The `DistribMulActionHom` underlying a `LinearMap`. -/ def toDistribMulActionHom (f : M →ₛₗ[σ] M₃) : DistribMulActionHom σ.toMonoidHom M M₃ := { f with map_zero' := show f 0 = 0 from map_zero f } diff --git a/Mathlib/Algebra/MvPolynomial/Basic.lean b/Mathlib/Algebra/MvPolynomial/Basic.lean index c5e9e3036cee4..de4f2a12c09bb 100644 --- a/Mathlib/Algebra/MvPolynomial/Basic.lean +++ b/Mathlib/Algebra/MvPolynomial/Basic.lean @@ -509,8 +509,6 @@ section Coeff /-- The coefficient of the monomial `m` in the multi-variable polynomial `p`. -/ def coeff (m : σ →₀ ℕ) (p : MvPolynomial σ R) : R := @DFunLike.coe ((σ →₀ ℕ) →₀ R) _ _ _ p m - -- Porting note: I changed this from `@CoeFun.coe _ _ (MonoidAlgebra.coeFun _ _) p m` because - -- I think it should work better syntactically. They are defeq. @[simp] theorem mem_support_iff {p : MvPolynomial σ R} {m : σ →₀ ℕ} : m ∈ p.support ↔ p.coeff m ≠ 0 := by diff --git a/Mathlib/Analysis/InnerProductSpace/PiL2.lean b/Mathlib/Analysis/InnerProductSpace/PiL2.lean index 5fb4aa1c20293..2c14b91ede9f1 100644 --- a/Mathlib/Analysis/InnerProductSpace/PiL2.lean +++ b/Mathlib/Analysis/InnerProductSpace/PiL2.lean @@ -306,7 +306,6 @@ theorem repr_injective : cases g congr --- Porting note: `CoeFun` → `FunLike` /-- `b i` is the `i`th basis vector. -/ instance instFunLike : FunLike (OrthonormalBasis ι 𝕜 E) ι E where coe b i := by classical exact b.repr.symm (EuclideanSpace.single i (1 : 𝕜)) diff --git a/Mathlib/LinearAlgebra/Multilinear/Basic.lean b/Mathlib/LinearAlgebra/Multilinear/Basic.lean index 21a0244972b27..5fa2ebd1c7ff4 100644 --- a/Mathlib/LinearAlgebra/Multilinear/Basic.lean +++ b/Mathlib/LinearAlgebra/Multilinear/Basic.lean @@ -106,7 +106,6 @@ variable [Semiring R] [∀ i, AddCommMonoid (M i)] [∀ i, AddCommMonoid (M₁ i [AddCommMonoid M₃] [AddCommMonoid M'] [∀ i, Module R (M i)] [∀ i, Module R (M₁ i)] [Module R M₂] [Module R M₃] [Module R M'] (f f' : MultilinearMap R M₁ M₂) --- Porting note: Replaced CoeFun with FunLike instance instance : FunLike (MultilinearMap R M₁ M₂) (∀ i, M₁ i) M₂ where coe f := f.toFun coe_injective' f g h := by cases f; cases g; cases h; rfl diff --git a/Mathlib/NumberTheory/ArithmeticFunction.lean b/Mathlib/NumberTheory/ArithmeticFunction.lean index 9e6e2e67edaef..71c09ddd1baad 100644 --- a/Mathlib/NumberTheory/ArithmeticFunction.lean +++ b/Mathlib/NumberTheory/ArithmeticFunction.lean @@ -92,7 +92,6 @@ section Zero variable [Zero R] --- porting note: used to be `CoeFun` instance : FunLike (ArithmeticFunction R) ℕ R := inferInstanceAs (FunLike (ZeroHom ℕ R) ℕ R) diff --git a/Mathlib/Order/Heyting/Hom.lean b/Mathlib/Order/Heyting/Hom.lean index 6b0940817f6d3..ab4677d3ab3cb 100644 --- a/Mathlib/Order/Heyting/Hom.lean +++ b/Mathlib/Order/Heyting/Hom.lean @@ -246,13 +246,6 @@ instance instHeytingHomClass : HeytingHomClass (HeytingHom α β) α β where map_bot f := f.map_bot' map_himp := HeytingHom.map_himp' - --- Porting note: CoeFun undesired here in lean 4 --- /-- Helper instance for when there's too many metavariables to apply `DFunLike.CoeFun` --- directly. -/ --- instance : CoeFun (HeytingHom α β) fun _ => α → β := --- DFunLike.hasCoeToFun - -- @[simp] -- Porting note: not in simp-nf, simp can simplify lhs. Added aux simp lemma theorem toFun_eq_coe {f : HeytingHom α β} : f.toFun = ⇑f := rfl diff --git a/Mathlib/Order/Hom/CompleteLattice.lean b/Mathlib/Order/Hom/CompleteLattice.lean index 8fecfd2875b9b..125753699728c 100644 --- a/Mathlib/Order/Hom/CompleteLattice.lean +++ b/Mathlib/Order/Hom/CompleteLattice.lean @@ -592,12 +592,6 @@ def tosSupHom (f : CompleteLatticeHom α β) : sSupHom α β := def toBoundedLatticeHom (f : CompleteLatticeHom α β) : BoundedLatticeHom α β := f --- Porting note: We do not want CoeFun for this in lean 4 --- /-- Helper instance for when there's too many metavariables to apply `fun_like.has_coe_toFun` --- directly. -/ --- instance : CoeFun (CompleteLatticeHom α β) fun _ => α → β := --- DFunLike.hasCoeToFun - lemma toFun_eq_coe (f : CompleteLatticeHom α β) : f.toFun = f := rfl @[simp] lemma coe_tosInfHom (f : CompleteLatticeHom α β) : ⇑f.tosInfHom = f := rfl diff --git a/Mathlib/RingTheory/RingInvo.lean b/Mathlib/RingTheory/RingInvo.lean index 0ebf763f2a467..86cbcf2eb18f4 100644 --- a/Mathlib/RingTheory/RingInvo.lean +++ b/Mathlib/RingTheory/RingInvo.lean @@ -83,10 +83,6 @@ def mk' (f : R →+* Rᵐᵒᵖ) (involution : ∀ r, (f (f r).unop).unop = r) : right_inv := fun _ => MulOpposite.unop_injective <| involution _ involution' := involution } --- Porting note: removed CoeFun instance, undesired in lean4 --- instance : CoeFun (RingInvo R) fun _ => R → Rᵐᵒᵖ := --- ⟨fun f => f.toRingEquiv.toFun⟩ - @[simp] theorem involution (f : RingInvo R) (x : R) : (f (f x).unop).unop = x := f.involution' x diff --git a/Mathlib/Topology/Algebra/Module/Basic.lean b/Mathlib/Topology/Algebra/Module/Basic.lean index 451dff66e30df..df022af891410 100644 --- a/Mathlib/Topology/Algebra/Module/Basic.lean +++ b/Mathlib/Topology/Algebra/Module/Basic.lean @@ -378,10 +378,6 @@ instance continuousSemilinearMapClass : map_continuous f := f.2 map_smulₛₗ f := f.toLinearMap.map_smul' --- see Note [function coercion] -/-- Coerce continuous linear maps to functions. -/ ---instance toFun' : CoeFun (M₁ →SL[σ₁₂] M₂) fun _ => M₁ → M₂ := ⟨DFunLike.coe⟩ - -- porting note (#10618): was `simp`, now `simp only` proves it theorem coe_mk (f : M₁ →ₛₗ[σ₁₂] M₂) (h) : (mk f h : M₁ →ₛₗ[σ₁₂] M₂) = f := rfl @@ -1634,11 +1630,6 @@ instance continuousSemilinearEquivClass : map_continuous := continuous_toFun inv_continuous := continuous_invFun --- see Note [function coercion] --- /-- Coerce continuous linear equivs to maps. -/ --- instance : CoeFun (M₁ ≃SL[σ₁₂] M₂) fun _ => M₁ → M₂ := --- ⟨fun f => f⟩ - theorem coe_apply (e : M₁ ≃SL[σ₁₂] M₂) (b : M₁) : (e : M₁ →SL[σ₁₂] M₂) b = e b := rfl diff --git a/Mathlib/Topology/Category/TopCat/Basic.lean b/Mathlib/Topology/Category/TopCat/Basic.lean index 468168d5c7fb5..8cf13a2f355f4 100644 --- a/Mathlib/Topology/Category/TopCat/Basic.lean +++ b/Mathlib/Topology/Category/TopCat/Basic.lean @@ -47,12 +47,6 @@ instance : CoeSort TopCat Type* where instance topologicalSpaceUnbundled (X : TopCat) : TopologicalSpace X := X.str --- We leave this temporarily as a reminder of the downstream instances #13170 --- -- Porting note: cannot find a coercion to function otherwise --- -- attribute [instance] ConcreteCategory.instFunLike in --- instance (X Y : TopCat.{u}) : CoeFun (X ⟶ Y) fun _ => X → Y where --- coe (f : C(X, Y)) := f - instance instFunLike (X Y : TopCat) : FunLike (X ⟶ Y) X Y := inferInstanceAs <| FunLike C(X, Y) X Y From d0df6356048affd1a0c86f935b0d04bf36a402d4 Mon Sep 17 00:00:00 2001 From: Yakov Pechersky Date: Thu, 17 Oct 2024 15:03:58 +0000 Subject: [PATCH 085/505] feat(QuotientGroup,Ideal/Quotient): group is finite if quotient is finite (#15425) Does not actually require the subgroup to be normal Also provides a breakdown of the group as the disjoint indexed union of cosets Co-authored-by: Yakov Pechersky --- Mathlib/GroupTheory/Coset/Basic.lean | 21 +++++++++++++++++ Mathlib/GroupTheory/GroupAction/Basic.lean | 23 +++++++++++++++++++ Mathlib/GroupTheory/QuotientGroup/Finite.lean | 5 ++++ Mathlib/RingTheory/Ideal/Quotient.lean | 11 +++++++++ 4 files changed, 60 insertions(+) diff --git a/Mathlib/GroupTheory/Coset/Basic.lean b/Mathlib/GroupTheory/Coset/Basic.lean index d9ede0a36ed67..ab5461c4022dc 100644 --- a/Mathlib/GroupTheory/Coset/Basic.lean +++ b/Mathlib/GroupTheory/Coset/Basic.lean @@ -458,6 +458,18 @@ theorem preimage_image_mk_eq_mul (N : Subgroup α) (s : Set α) : rw [preimage_image_mk_eq_iUnion_image, iUnion_subtype, ← image2_mul, ← iUnion_image_right] simp only [SetLike.mem_coe] +open MulAction in +@[to_additive] +lemma orbit_mk_eq_smul (x : α) : MulAction.orbitRel.Quotient.orbit (x : α ⧸ s) = x • s := by + ext + rw [orbitRel.Quotient.mem_orbit] + simpa [mem_smul_set_iff_inv_smul_mem, ← leftRel_apply] using Setoid.comm' _ + +@[to_additive] +lemma orbit_eq_out'_smul (x : α ⧸ s) : MulAction.orbitRel.Quotient.orbit x = x.out' • s := by + induction x using QuotientGroup.induction_on + simp only [orbit_mk_eq_smul, ← eq_class_eq_leftCoset, Quotient.out_eq'] + end QuotientGroup namespace Subgroup @@ -706,4 +718,13 @@ noncomputable def preimageMkEquivSubgroupProdSet (s : Subgroup α) (t : Set (α left_inv := fun ⟨a, _⟩ => Subtype.eq <| show _ * _ = a by simp right_inv := fun ⟨⟨a, ha⟩, ⟨x, hx⟩⟩ => by ext <;> simp [ha] +open MulAction in +/-- A group is made up of a disjoint union of cosets of a subgroup. -/ +@[to_additive "An additive group is made up of a disjoint union of cosets of an additive +subgroup."] +lemma univ_eq_iUnion_smul (H : Subgroup α) : + (Set.univ (α := α)) = ⋃ x : α ⧸ H, x.out' • (H : Set _) := by + simp_rw [univ_eq_iUnion_orbit H.op, orbit_eq_out'_smul] + rfl + end QuotientGroup diff --git a/Mathlib/GroupTheory/GroupAction/Basic.lean b/Mathlib/GroupTheory/GroupAction/Basic.lean index cc847556ec7f8..9f99144a7769a 100644 --- a/Mathlib/GroupTheory/GroupAction/Basic.lean +++ b/Mathlib/GroupTheory/GroupAction/Basic.lean @@ -108,6 +108,10 @@ lemma snd_mem_orbit_of_mem_orbit {x y : α × β} (h : x ∈ MulAction.orbit M y rcases h with ⟨g, rfl⟩ exact mem_orbit _ _ +@[to_additive] +lemma _root_.Finite.finite_mulAction_orbit [Finite M] (a : α) : Set.Finite (orbit M a) := + Set.finite_range _ + variable (M) @[to_additive] @@ -625,6 +629,25 @@ def selfEquivSigmaOrbits : α ≃ Σω : Ω, orbit G ω.out' := Equiv.sigmaCongrRight fun _ => Equiv.Set.ofEq <| orbitRel.Quotient.orbit_eq_orbit_out _ Quotient.out_eq' +/-- Decomposition of a type `X` as a disjoint union of its orbits under a group action. +Phrased as a set union. See `MulAction.selfEquivSigmaOrbits` for the type isomorphism. -/ +@[to_additive "Decomposition of a type `X` as a disjoint union of its orbits under an additive group +action. Phrased as a set union. See `AddAction.selfEquivSigmaOrbits` for the type isomorphism."] +lemma univ_eq_iUnion_orbit : + Set.univ (α := α) = ⋃ x : Ω, x.orbit := by + ext x + simp only [Set.mem_univ, Set.mem_iUnion, true_iff] + exact ⟨Quotient.mk'' x, by simp⟩ + +@[to_additive] +lemma _root_.Finite.of_finite_mulAction_orbitRel_quotient [Finite G] [Finite Ω] : Finite α := by + rw [(selfEquivSigmaOrbits' G _).finite_iff] + have : ∀ g : Ω, Finite g.orbit := by + intro g + induction g using Quotient.inductionOn' + simpa [Set.finite_coe_iff] using Finite.finite_mulAction_orbit _ + exact Finite.instSigma + variable (β) @[to_additive] diff --git a/Mathlib/GroupTheory/QuotientGroup/Finite.lean b/Mathlib/GroupTheory/QuotientGroup/Finite.lean index 76877f51f87b7..6d234fdd50c16 100644 --- a/Mathlib/GroupTheory/QuotientGroup/Finite.lean +++ b/Mathlib/GroupTheory/QuotientGroup/Finite.lean @@ -50,4 +50,9 @@ noncomputable def fintypeOfKerOfCodom [Fintype g.ker] : Fintype G := noncomputable def fintypeOfDomOfCoker [Normal f.range] [Fintype <| G ⧸ f.range] : Fintype G := fintypeOfKerLeRange _ (mk' f.range) fun x => (eq_one_iff x).mp +@[to_additive] +lemma _root_.Finite.of_finite_quot_finite_subgroup {H : Subgroup G} [Finite H] [Finite (G ⧸ H)] : + Finite G := + Finite.of_equiv _ (groupEquivQuotientProdSubgroup (s := H)).symm + end Group diff --git a/Mathlib/RingTheory/Ideal/Quotient.lean b/Mathlib/RingTheory/Ideal/Quotient.lean index 3889f7377adb0..182a09a1dac95 100644 --- a/Mathlib/RingTheory/Ideal/Quotient.lean +++ b/Mathlib/RingTheory/Ideal/Quotient.lean @@ -3,6 +3,7 @@ Copyright (c) 2018 Kenny Lau. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau, Chris Hughes, Mario Carneiro, Anne Baanen -/ +import Mathlib.GroupTheory.QuotientGroup.Finite import Mathlib.LinearAlgebra.Quotient import Mathlib.RingTheory.Congruence.Basic import Mathlib.RingTheory.Ideal.Basic @@ -348,4 +349,14 @@ theorem map_pi {ι : Type*} [Finite ι] {ι' : Type w} (x : ι → R) (hi : ∀ end Pi +open scoped Pointwise in +/-- A ring is made up of a disjoint union of cosets of an ideal. -/ +lemma univ_eq_iUnion_image_add {R : Type*} [Ring R] (I : Ideal R) : + (Set.univ (α := R)) = ⋃ x : R ⧸ I, x.out' +ᵥ (I : Set R) := + QuotientAddGroup.univ_eq_iUnion_vadd I.toAddSubgroup + +lemma _root_.Finite.of_finite_quot_finite_ideal {R : Type*} [Ring R] {I : Ideal R} + [hI : Finite I] [h : Finite (R ⧸ I)] : Finite R := + @Finite.of_finite_quot_finite_addSubgroup _ _ _ hI h + end Ideal From 0b86556786ec68a8d3c198f66952df4d3b771cb2 Mon Sep 17 00:00:00 2001 From: FR Date: Thu, 17 Oct 2024 15:03:59 +0000 Subject: [PATCH 086/505] chore(SetTheory/Game/Basic): fix docs (#17121) --- Mathlib/SetTheory/Game/Basic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/SetTheory/Game/Basic.lean b/Mathlib/SetTheory/Game/Basic.lean index 13900670705bd..3504ce49ac81a 100644 --- a/Mathlib/SetTheory/Game/Basic.lean +++ b/Mathlib/SetTheory/Game/Basic.lean @@ -247,7 +247,7 @@ Hence we define them here. -/ /-- The product of `x = {xL | xR}` and `y = {yL | yR}` is -`{xL*y + x*yL - xL*yL, xR*y + x*yR - xR*yR | xL*y + x*yR - xL*yR, x*yL + xR*y - xR*yL }`. -/ +`{xL*y + x*yL - xL*yL, xR*y + x*yR - xR*yR | xL*y + x*yR - xL*yR, xR*y + x*yL - xR*yL}`. -/ instance : Mul PGame.{u} := ⟨fun x y => by induction x generalizing y with | mk xl xr _ _ IHxl IHxr => _ From cff08d545656d8cabf918cdcbde50f4f6105ae0b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 17 Oct 2024 15:04:01 +0000 Subject: [PATCH 087/505] chore(Order/Hom): let `gcongr` know about `OrderEmbedding` and `OrderIso` (#17224) From LeanAPAP --- Mathlib/NumberTheory/NumberField/Discriminant.lean | 2 +- Mathlib/Order/Hom/Basic.lean | 7 +++++++ 2 files changed, 8 insertions(+), 1 deletion(-) diff --git a/Mathlib/NumberTheory/NumberField/Discriminant.lean b/Mathlib/NumberTheory/NumberField/Discriminant.lean index 758d8a99b48ed..07199e222dcb4 100644 --- a/Mathlib/NumberTheory/NumberField/Discriminant.lean +++ b/Mathlib/NumberTheory/NumberField/Discriminant.lean @@ -308,7 +308,7 @@ theorem minkowskiBound_lt_boundOfDiscBdd : minkowskiBound K ↑1 < boundOfDiscBd coe_mul, ENNReal.coe_pow, coe_ofNat, show sqrt N = (1 : ℝ≥0∞) * sqrt N by rw [one_mul]] gcongr · exact pow_le_one₀ (by positivity) (by norm_num) - · rwa [sqrt_le_sqrt, ← NNReal.coe_le_coe, coe_nnnorm, Int.norm_eq_abs, ← Int.cast_abs, + · rwa [← NNReal.coe_le_coe, coe_nnnorm, Int.norm_eq_abs, ← Int.cast_abs, NNReal.coe_natCast, ← Int.cast_natCast, Int.cast_le] · exact one_le_two · exact rank_le_rankOfDiscrBdd hK diff --git a/Mathlib/Order/Hom/Basic.lean b/Mathlib/Order/Hom/Basic.lean index 96b44c3f6a264..c9ab0123359da 100644 --- a/Mathlib/Order/Hom/Basic.lean +++ b/Mathlib/Order/Hom/Basic.lean @@ -148,6 +148,9 @@ protected theorem monotone (f : F) : Monotone f := fun _ _ => map_rel f protected theorem mono (f : F) : Monotone f := fun _ _ => map_rel f +@[gcongr] protected lemma GCongr.mono (f : F) {a b : α} (hab : a ≤ b) : f a ≤ f b := + OrderHomClass.mono f hab + /-- Turn an element of a type `F` satisfying `OrderHomClass F α β` into an actual `OrderHom`. This is declared as the default coercion from `F` to `α →o β`. -/ @[coe] @@ -921,6 +924,8 @@ variable [LE α] [LE β] theorem le_iff_le (e : α ≃o β) {x y : α} : e x ≤ e y ↔ x ≤ y := e.map_rel_iff +@[gcongr] protected alias ⟨_, GCongr.orderIso_apply_le_apply⟩ := le_iff_le + theorem le_symm_apply (e : α ≃o β) {x : α} {y : β} : x ≤ e.symm y ↔ e x ≤ y := e.rel_symm_apply @@ -941,6 +946,8 @@ protected theorem strictMono (e : α ≃o β) : StrictMono e := theorem lt_iff_lt (e : α ≃o β) {x y : α} : e x < e y ↔ x < y := e.toOrderEmbedding.lt_iff_lt +@[gcongr] protected alias ⟨_, GCongr.orderIso_apply_lt_apply⟩ := lt_iff_lt + /-- Converts an `OrderIso` into a `RelIso (<) (<)`. -/ def toRelIsoLT (e : α ≃o β) : ((· < ·) : α → α → Prop) ≃r ((· < ·) : β → β → Prop) := ⟨e.toEquiv, lt_iff_lt e⟩ From cc8cf2ffff4f61b00d30cd82f66d1728e95dc4c3 Mon Sep 17 00:00:00 2001 From: Mario Carneiro Date: Thu, 17 Oct 2024 15:04:03 +0000 Subject: [PATCH 088/505] feat: add `--force` flag to shake (#17453) to disable the `lake build --no-build` step. There have been [reports](https://leanprover.zulipchat.com/#narrow/stream/113488-general/topic/.60lake.20build.20--no-build.60.20at.20.60v4.2E13.2E0-rc3.60.20is.20broken.3F/near/474875999) that this step is somewhat flaky, so this option gives people the ability to skip it without turning off `shake` entirely. Co-authored-by: Mario Carneiro --- Shake/Main.lean | 13 ++++++++++--- 1 file changed, 10 insertions(+), 3 deletions(-) diff --git a/Shake/Main.lean b/Shake/Main.lean index e5b7645a31d81..20a8f24e116f0 100644 --- a/Shake/Main.lean +++ b/Shake/Main.lean @@ -35,6 +35,9 @@ Arguments: provided module(s) will be checked. Options: + --force + Skips the `lake build --no-build` sanity check + --fix Apply the suggested fixes directly. Make sure you have a clean checkout before running this, so you can review the changes. @@ -380,6 +383,8 @@ def toBitset (s : State) (ns : List Name) : Bitset := structure Args where /-- `--help`: shows the help -/ help : Bool := false + /-- `--force`: skips the `lake build --no-build` sanity check -/ + force : Bool := false /-- `--no-downstream`: disables downstream mode -/ downstream : Bool := true /-- `--gh-style`: output messages that can be parsed by `gh-problem-matcher-wrap` -/ @@ -422,6 +427,7 @@ def main (args : List String) : IO UInt32 := do let rec parseArgs (args : Args) : List String → Args | [] => args | "--help" :: rest => parseArgs { args with help := true } rest + | "--force" :: rest => parseArgs { args with force := true } rest | "--no-downstream" :: rest => parseArgs { args with downstream := false } rest | "--fix" :: rest => parseArgs { args with fix := true } rest | "--explain" :: rest => parseArgs { args with explain := true } rest @@ -438,9 +444,10 @@ def main (args : List String) : IO UInt32 := do IO.println help IO.Process.exit 0 - if (← IO.Process.output { cmd := "lake", args := #["build", "--no-build"] }).exitCode != 0 then - IO.println "There are out of date oleans. Run `lake build` or `lake exe cache get` first" - IO.Process.exit 1 + if !args.force then + if (← IO.Process.output { cmd := "lake", args := #["build", "--no-build"] }).exitCode != 0 then + IO.println "There are out of date oleans. Run `lake build` or `lake exe cache get` first" + IO.Process.exit 1 -- Parse the `--cfg` argument let srcSearchPath ← initSrcSearchPath From 10e0d7deb5528d3e0eb8331b42f15f24b0559767 Mon Sep 17 00:00:00 2001 From: Anatole Dedecker Date: Thu, 17 Oct 2024 15:04:04 +0000 Subject: [PATCH 089/505] feat: seminorms for the product topology (#17493) --- Mathlib/Analysis/LocallyConvex/WithSeminorms.lean | 13 +++++++++++-- 1 file changed, 11 insertions(+), 2 deletions(-) diff --git a/Mathlib/Analysis/LocallyConvex/WithSeminorms.lean b/Mathlib/Analysis/LocallyConvex/WithSeminorms.lean index dff476482e167..5238beff83992 100644 --- a/Mathlib/Analysis/LocallyConvex/WithSeminorms.lean +++ b/Mathlib/Analysis/LocallyConvex/WithSeminorms.lean @@ -877,13 +877,22 @@ protected def SeminormFamily.sigma {κ : ι → Type*} (p : (i : ι) → Seminor theorem withSeminorms_iInf {κ : ι → Type*} [Nonempty ((i : ι) × κ i)] [∀ i, Nonempty (κ i)] {p : (i : ι) → SeminormFamily 𝕜 E (κ i)} {t : ι → TopologicalSpace E} - [∀ i, @TopologicalAddGroup E (t i) _] (hp : ∀ i, WithSeminorms (topology := t i) (p i)) : + (hp : ∀ i, WithSeminorms (topology := t i) (p i)) : WithSeminorms (topology := ⨅ i, t i) (SeminormFamily.sigma p) := by - have : @TopologicalAddGroup E (⨅ i, t i) _ := topologicalAddGroup_iInf (fun i ↦ inferInstance) + have : ∀ i, @TopologicalAddGroup E (t i) _ := + fun i ↦ @WithSeminorms.topologicalAddGroup _ _ _ _ _ _ _ (t i) _ (hp i) + have : @TopologicalAddGroup E (⨅ i, t i) _ := topologicalAddGroup_iInf inferInstance simp_rw [@SeminormFamily.withSeminorms_iff_topologicalSpace_eq_iInf _ _ _ _ _ _ _ (_)] at hp ⊢ rw [iInf_sigma] exact iInf_congr hp +theorem withSeminorms_pi {κ : ι → Type*} {E : ι → Type*} + [∀ i, AddCommGroup (E i)] [∀ i, Module 𝕜 (E i)] [∀ i, TopologicalSpace (E i)] + [Nonempty ((i : ι) × κ i)] [∀ i, Nonempty (κ i)] {p : (i : ι) → SeminormFamily 𝕜 (E i) (κ i)} + (hp : ∀ i, WithSeminorms (p i)) : + WithSeminorms (SeminormFamily.sigma (fun i ↦ (p i).comp (LinearMap.proj i))) := + withSeminorms_iInf fun i ↦ (LinearMap.proj i).withSeminorms_induced (hp i) + end TopologicalConstructions section TopologicalProperties From 9dd0528bdc2e9987bc51e623517606202d56f143 Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Thu, 17 Oct 2024 15:04:05 +0000 Subject: [PATCH 090/505] feat(GroupTheory/Coset/Basic): products, `leftRel` and `rightRel` (#17585) Add lemmas relating `leftRel` and `rightRel` for a (pairwise or indexed) product of subgroups to the corresponding products of setoids. From AperiodicMonotilesLean. --- Mathlib/GroupTheory/Coset/Basic.lean | 28 ++++++++++++++++++++++++++++ 1 file changed, 28 insertions(+) diff --git a/Mathlib/GroupTheory/Coset/Basic.lean b/Mathlib/GroupTheory/Coset/Basic.lean index ab5461c4022dc..e7aa809a354ae 100644 --- a/Mathlib/GroupTheory/Coset/Basic.lean +++ b/Mathlib/GroupTheory/Coset/Basic.lean @@ -262,6 +262,20 @@ instance leftRelDecidable [DecidablePred (· ∈ s)] : DecidableRel (leftRel s). rw [leftRel_eq] exact ‹DecidablePred (· ∈ s)› _ +@[to_additive] +lemma leftRel_prod {β : Type*} [Group β] (s' : Subgroup β) : + leftRel (s.prod s') = (leftRel s).prod (leftRel s') := by + refine Setoid.ext fun x y ↦ ?_ + rw [Setoid.prod_apply] + simp_rw [leftRel_apply] + rfl + +@[to_additive] +lemma leftRel_pi {ι : Type*} {β : ι → Type*} [∀ i, Group (β i)] (s' : ∀ i, Subgroup (β i)) : + leftRel (Subgroup.pi Set.univ s') = @piSetoid _ _ fun i ↦ leftRel (s' i) := by + refine Setoid.ext fun x y ↦ ?_ + simp [Setoid.piSetoid_apply, leftRel_apply, Subgroup.mem_pi] + /-- `α ⧸ s` is the quotient type representing the left cosets of `s`. If `s` is a normal subgroup, `α ⧸ s` is a group -/ @[to_additive "`α ⧸ s` is the quotient type representing the left cosets of `s`. If `s` is a normal @@ -304,6 +318,20 @@ instance rightRelDecidable [DecidablePred (· ∈ s)] : DecidableRel (rightRel s rw [rightRel_eq] exact ‹DecidablePred (· ∈ s)› _ +@[to_additive] +lemma rightRel_prod {β : Type*} [Group β] (s' : Subgroup β) : + rightRel (s.prod s') = (rightRel s).prod (rightRel s') := by + refine Setoid.ext fun x y ↦ ?_ + rw [Setoid.prod_apply] + simp_rw [rightRel_apply] + rfl + +@[to_additive] +lemma rightRel_pi {ι : Type*} {β : ι → Type*} [∀ i, Group (β i)] (s' : ∀ i, Subgroup (β i)) : + rightRel (Subgroup.pi Set.univ s') = @piSetoid _ _ fun i ↦ rightRel (s' i) := by + refine Setoid.ext fun x y ↦ ?_ + simp [Setoid.piSetoid_apply, rightRel_apply, Subgroup.mem_pi] + /-- Right cosets are in bijection with left cosets. -/ @[to_additive "Right cosets are in bijection with left cosets."] def quotientRightRelEquivQuotientLeftRel : Quotient (QuotientGroup.rightRel s) ≃ α ⧸ s where From f66852471cf4c2c6cb3b707dc1e9be6929360662 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 17 Oct 2024 15:04:07 +0000 Subject: [PATCH 091/505] feat: translating the graph of a homomorphism (#17647) From PFR --- Mathlib/Data/Set/Pointwise/SMul.lean | 26 ++++++++++++++++++++++++++ 1 file changed, 26 insertions(+) diff --git a/Mathlib/Data/Set/Pointwise/SMul.lean b/Mathlib/Data/Set/Pointwise/SMul.lean index e3ed79c944b46..00fe29ebfea63 100644 --- a/Mathlib/Data/Set/Pointwise/SMul.lean +++ b/Mathlib/Data/Set/Pointwise/SMul.lean @@ -406,6 +406,32 @@ lemma smul_set_disjoint_iff : Disjoint (a • s) (a • t) ↔ Disjoint s t := b end Group +section Group +variable [Group α] [CommGroup β] [FunLike F α β] [MonoidHomClass F α β] + +@[to_additive] +lemma smul_graphOn (x : α × β) (s : Set α) (f : F) : + x • s.graphOn f = (x.1 • s).graphOn fun a ↦ x.2 / f x.1 * f a := by + ext ⟨a, b⟩ + simp [mem_smul_set_iff_inv_smul_mem, Prod.ext_iff, and_comm (a := _ = a), inv_mul_eq_iff_eq_mul, + mul_left_comm _ _⁻¹, eq_inv_mul_iff_mul_eq, ← mul_div_right_comm, div_eq_iff_eq_mul, mul_comm b] + +@[to_additive] +lemma smul_graphOn_univ (x : α × β) (f : F) : + x • univ.graphOn f = univ.graphOn fun a ↦ x.2 / f x.1 * f a := by simp [smul_graphOn] + +end Group + +section CommGroup +variable [CommGroup α] + +@[to_additive] lemma smul_div_smul_comm (a : α) (s : Set α) (b : α) (t : Set α) : + a • s / b • t = (a / b) • (s / t) := by + simp_rw [← image_smul, smul_eq_mul, ← singleton_mul, mul_div_mul_comm _ s, + singleton_div_singleton] + +end CommGroup + section GroupWithZero variable [GroupWithZero α] [MulAction α β] {s t : Set β} {a : α} From 1cfdb12acbfe9bb8e0c42aa6a32029b0ecb80b84 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 17 Oct 2024 15:04:08 +0000 Subject: [PATCH 092/505] feat: Over a finite ring, a module is finite iff it is finite dimensional (#17707) From PFR --- Mathlib/FieldTheory/AxGrothendieck.lean | 3 +-- Mathlib/FieldTheory/PrimitiveElement.lean | 2 +- .../LinearAlgebra/FiniteDimensional/Defs.lean | 5 ----- Mathlib/RingTheory/Finiteness.lean | 22 +++++++++++++++++++ .../LocalRing/ResidueField/Basic.lean | 6 ++--- 5 files changed, 27 insertions(+), 11 deletions(-) diff --git a/Mathlib/FieldTheory/AxGrothendieck.lean b/Mathlib/FieldTheory/AxGrothendieck.lean index e386e095dad39..9ddfa55c77de3 100644 --- a/Mathlib/FieldTheory/AxGrothendieck.lean +++ b/Mathlib/FieldTheory/AxGrothendieck.lean @@ -48,8 +48,7 @@ theorem ax_grothendieck_of_locally_finite {ι K R : Type*} [Field K] [Finite K] (mem_union_left _ (mem_biUnion.2 ⟨i, mem_univ _, mem_image_of_mem _ hk⟩)) letI := isNoetherian_adjoin_finset s fun x _ => Algebra.IsIntegral.isIntegral (R := K) x letI := Module.IsNoetherian.finite K (Algebra.adjoin K (s : Set R)) - letI : Finite (Algebra.adjoin K (s : Set R)) := - FiniteDimensional.finite_of_finite K (Algebra.adjoin K (s : Set R)) + letI : Finite (Algebra.adjoin K (s : Set R)) := Module.finite_of_finite K -- The restriction of the polynomial map, `ps`, to the subalgebra generated by `s` let res : (ι → Algebra.adjoin K (s : Set R)) → ι → Algebra.adjoin K (s : Set R) := fun x i => ⟨eval (fun j : ι => (x j : R)) (ps i), eval_mem (hs₁ _) fun i => (x i).2⟩ diff --git a/Mathlib/FieldTheory/PrimitiveElement.lean b/Mathlib/FieldTheory/PrimitiveElement.lean index da215b5d8dbf4..09b73357bf4f2 100644 --- a/Mathlib/FieldTheory/PrimitiveElement.lean +++ b/Mathlib/FieldTheory/PrimitiveElement.lean @@ -63,7 +63,7 @@ theorem exists_primitive_element_of_finite_top [Finite E] : ∃ α : E, F⟮α /-- Primitive element theorem for finite dimensional extension of a finite field. -/ theorem exists_primitive_element_of_finite_bot [Finite F] [FiniteDimensional F E] : ∃ α : E, F⟮α⟯ = ⊤ := - haveI : Finite E := FiniteDimensional.finite_of_finite F E + haveI : Finite E := Module.finite_of_finite F exists_primitive_element_of_finite_top F E end PrimitiveElementFinite diff --git a/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean b/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean index 2bce02c0a50fc..14b5f34c2a2b3 100644 --- a/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean +++ b/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean @@ -107,11 +107,6 @@ instance finiteDimensional_pi' {ι : Type*} [Finite ι] (M : ι → Type*) [∀ noncomputable def fintypeOfFintype [Fintype K] [FiniteDimensional K V] : Fintype V := Module.fintypeOfFintype (@finsetBasis K V _ _ _ (iff_fg.2 inferInstance)) -theorem finite_of_finite [Finite K] [FiniteDimensional K V] : Finite V := by - cases nonempty_fintype K - haveI := fintypeOfFintype K V - infer_instance - variable {K V} /-- If a vector space has a finite basis, then it is finite-dimensional. -/ diff --git a/Mathlib/RingTheory/Finiteness.lean b/Mathlib/RingTheory/Finiteness.lean index 1425279d2a674..264b1bb8d9350 100644 --- a/Mathlib/RingTheory/Finiteness.lean +++ b/Mathlib/RingTheory/Finiteness.lean @@ -537,6 +537,22 @@ lemma exists_fin' [Module.Finite R M] : ∃ (n : ℕ) (f : (Fin n → R) →ₗ[ refine ⟨n, Basis.constr (Pi.basisFun R _) ℕ s, ?_⟩ rw [← LinearMap.range_eq_top, Basis.constr_range, hs] +variable (R) in +lemma _root_.Module.finite_of_finite [Finite R] [Module.Finite R M] : Finite M := by + obtain ⟨n, f, hf⟩ := exists_fin' R M; exact .of_surjective f hf + +@[deprecated (since := "2024-10-13")] +alias _root_.FiniteDimensional.finite_of_finite := finite_of_finite + +-- See note [lower instance priority] +instance (priority := 100) of_finite [Finite M] : Module.Finite R M := by + cases nonempty_fintype M + exact ⟨⟨Finset.univ, by rw [Finset.coe_univ]; exact Submodule.span_univ⟩⟩ + +/-- A module over a finite ring has finite dimension iff it is finite. -/ +lemma _root_.Module.finite_iff_finite [Finite R] : Module.Finite R M ↔ Finite M := + ⟨fun _ ↦ finite_of_finite R, fun _ ↦ .of_finite⟩ + theorem of_surjective [hM : Module.Finite R M] (f : M →ₗ[R] N) (hf : Surjective f) : Module.Finite R N := ⟨by @@ -620,6 +636,12 @@ instance span_singleton (x : M) : Module.Finite R (R ∙ x) := instance span_finset (s : Finset M) : Module.Finite R (span R (s : Set M)) := ⟨(Submodule.fg_top _).mpr ⟨s, rfl⟩⟩ +lemma _root_.Set.Finite.submoduleSpan [Finite R] {s : Set M} (hs : s.Finite) : + (Submodule.span R s : Set M).Finite := by + lift s to Finset M using hs + rw [Set.Finite, ← Module.finite_iff_finite (R := R)] + dsimp + infer_instance theorem Module.End.isNilpotent_iff_of_finite {R M : Type*} [CommSemiring R] [AddCommMonoid M] [Module R M] [Module.Finite R M] {f : End R M} : diff --git a/Mathlib/RingTheory/LocalRing/ResidueField/Basic.lean b/Mathlib/RingTheory/LocalRing/ResidueField/Basic.lean index 885d727ff53f4..ca82438849309 100644 --- a/Mathlib/RingTheory/LocalRing/ResidueField/Basic.lean +++ b/Mathlib/RingTheory/LocalRing/ResidueField/Basic.lean @@ -166,10 +166,10 @@ instance finiteDimensional_of_noetherian [IsNoetherian R S] : (LinearMap.range_eq_top.mpr Ideal.Quotient.mk_surjective) exact Algebra.algebra_ext _ _ (fun r => rfl) +-- We want to be able to refer to `hfin` +set_option linter.unusedVariables false in lemma finite_of_finite [IsNoetherian R S] (hfin : Finite (ResidueField R)) : - Finite (ResidueField S) := by - have := @finiteDimensional_of_noetherian R S - exact FiniteDimensional.finite_of_finite (ResidueField R) (ResidueField S) + Finite (ResidueField S) := Module.finite_of_finite (ResidueField R) end FiniteDimensional From 0d150bcbee58d7d46b6d72a2ee0f113c00e3ec34 Mon Sep 17 00:00:00 2001 From: Dagur Asgeirsson Date: Thu, 17 Oct 2024 15:04:10 +0000 Subject: [PATCH 093/505] chore(CategoryTheory/Sites): add some tests for sheafcompose instances (#17817) This PR adds a few tests to make sure that there are `HasSheafCompose` instances for the forgetful functor from modules to sets, for different size restrictions on the defining site. --- test/CategoryTheory/Sites/Whiskering.lean | 28 +++++++++++++++++++++++ 1 file changed, 28 insertions(+) create mode 100644 test/CategoryTheory/Sites/Whiskering.lean diff --git a/test/CategoryTheory/Sites/Whiskering.lean b/test/CategoryTheory/Sites/Whiskering.lean new file mode 100644 index 0000000000000..af2efea006a9d --- /dev/null +++ b/test/CategoryTheory/Sites/Whiskering.lean @@ -0,0 +1,28 @@ +import Mathlib.Algebra.Category.ModuleCat.Colimits +import Mathlib.Algebra.Category.ModuleCat.Limits +import Mathlib.Algebra.Category.ModuleCat.FilteredColimits +import Mathlib.CategoryTheory.Sites.Equivalence + +universe u + +open CategoryTheory GrothendieckTopology + +section Small + +variable {C : Type u} [SmallCategory C] (J : GrothendieckTopology C) (R : Type u) [Ring R] + +example : HasSheafCompose J (forget (ModuleCat.{u} R)) := inferInstance + +end Small + +section Large + +variable {C : Type (u+1)} [LargeCategory C] (J : GrothendieckTopology C) + +example (R : Type (u+1)) [Ring R] : HasSheafCompose J (forget (ModuleCat.{u+1} R)) := inferInstance + +variable [EssentiallySmall.{u} C] + +example (R : Type u) [Ring R] : HasSheafCompose J (forget (ModuleCat.{u} R)) := inferInstance + +end Large From c9b69692d65d8b44e452ca4692ed9b04eb5e463f Mon Sep 17 00:00:00 2001 From: Dagur Asgeirsson Date: Thu, 17 Oct 2024 15:04:11 +0000 Subject: [PATCH 094/505] chore(CategoryTheory/Sites): add some tests for preserving sheafification (#17818) This PR adds a few tests to make sure that there are `PreservesSheafification` instances for the forgetful functor from modules to sets, for different size restrictions on the defining site. --- .../Sites/PreservesSheafification.lean | 30 +++++++++++++++++++ 1 file changed, 30 insertions(+) create mode 100644 test/CategoryTheory/Sites/PreservesSheafification.lean diff --git a/test/CategoryTheory/Sites/PreservesSheafification.lean b/test/CategoryTheory/Sites/PreservesSheafification.lean new file mode 100644 index 0000000000000..813746eb759e5 --- /dev/null +++ b/test/CategoryTheory/Sites/PreservesSheafification.lean @@ -0,0 +1,30 @@ +import Mathlib.Algebra.Category.ModuleCat.Colimits +import Mathlib.Algebra.Category.ModuleCat.Limits +import Mathlib.Algebra.Category.ModuleCat.FilteredColimits +import Mathlib.CategoryTheory.Sites.Equivalence + +universe u + +open CategoryTheory GrothendieckTopology + +section Small + +variable {C : Type u} [SmallCategory C] (J : GrothendieckTopology C) (R : Type u) [Ring R] + +example : PreservesSheafification J (forget (ModuleCat.{u} R)) := inferInstance + +end Small + +section Large + +variable {C : Type (u+1)} [LargeCategory C] (J : GrothendieckTopology C) + +example (R : Type (u+1)) [Ring R] : PreservesSheafification J (forget (ModuleCat.{u+1} R)) := + inferInstance + +variable [EssentiallySmall.{u} C] + +example (R : Type u) [Ring R] : PreservesSheafification J (forget (ModuleCat.{u} R)) := + inferInstance + +end Large From a70b960b28617ef85de0b08c7cb24b3f92dcc964 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 17 Oct 2024 15:04:12 +0000 Subject: [PATCH 095/505] feat: the filtration of cylinder events (#17858) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The exterior σ-algebras of finite sets of `α` form a cofiltration indexed by `Finset α` From GibbsMeasure Co-authored-by: Kin Yau James Wong --- Mathlib/Probability/Process/Filtration.lean | 9 +++++++++ 1 file changed, 9 insertions(+) diff --git a/Mathlib/Probability/Process/Filtration.lean b/Mathlib/Probability/Process/Filtration.lean index cc95366031c1b..853bae3ead907 100644 --- a/Mathlib/Probability/Process/Filtration.lean +++ b/Mathlib/Probability/Process/Filtration.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Kexing Ying. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kexing Ying, Rémy Degenne -/ +import Mathlib.MeasureTheory.Constructions.Cylinders import Mathlib.MeasureTheory.Function.ConditionalExpectation.Real /-! @@ -328,6 +329,14 @@ alias memℒp_limitProcess_of_snorm_bdd := memℒp_limitProcess_of_eLpNorm_bdd end Limit +variable {α : Type*} + +/-- The exterior σ-algebras of finite sets of `α` form a cofiltration indexed by `Finset α`. -/ +def cylinderEventsCompl : Filtration (Finset α)ᵒᵈ (.pi (π := fun _ : α ↦ Ω)) where + seq Λ := cylinderEvents (↑(OrderDual.ofDual Λ))ᶜ + mono' _ _ h := cylinderEvents_mono <| Set.compl_subset_compl_of_subset h + le' _ := cylinderEvents_le_pi + end Filtration end MeasureTheory From 3f297f53816d8a85d03f8c08b3a7d475dfa51a6b Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Thu, 17 Oct 2024 15:04:13 +0000 Subject: [PATCH 096/505] chore: remove porting notes about adding/removing `@[ext]` attributes (#17873) Discussed with @jcommelin and @kim-em. All of these porting notes are about the change in behaviour of the `ext` tactic and attribute between Lean 3 and Lean 4: in Lean 4 it only accepts lemmas with conclusion `x = y` with `x` and `y` free variables, and only applies these if the types are reducibly equal. Since those behaviours are not likely to change in the future, we don't need to keep these porting notes. I propose that we should keep #11041 since it indicates where these `ext` lemmas have not yet been added. Closes: #5229 Closes: #11182 --- Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean | 1 - Mathlib/Algebra/Group/AddChar.lean | 1 - Mathlib/Algebra/Homology/HomologicalComplex.lean | 1 - Mathlib/Algebra/Polynomial/Laurent.lean | 1 - Mathlib/AlgebraicTopology/SimplexCategory.lean | 1 - Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean | 2 -- Mathlib/AlgebraicTopology/SplitSimplicialObject.lean | 1 - Mathlib/CategoryTheory/Abelian/Pseudoelements.lean | 6 ------ .../Bicategory/NaturalTransformation/Oplax.lean | 1 - Mathlib/CategoryTheory/Comma/Basic.lean | 2 -- Mathlib/CategoryTheory/Comma/StructuredArrow.lean | 4 ---- Mathlib/CategoryTheory/Elements.lean | 1 - Mathlib/CategoryTheory/Endomorphism.lean | 1 - Mathlib/CategoryTheory/Functor/Category.lean | 4 ---- Mathlib/CategoryTheory/Idempotents/Karoubi.lean | 1 - Mathlib/CategoryTheory/Iso.lean | 4 ---- Mathlib/CategoryTheory/Limits/Cones.lean | 2 -- Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean | 2 -- Mathlib/CategoryTheory/Monoidal/Bimod.lean | 1 - Mathlib/CategoryTheory/Monoidal/Braided/Basic.lean | 2 -- Mathlib/CategoryTheory/Monoidal/CommMon_.lean | 2 -- Mathlib/CategoryTheory/Monoidal/Mod_.lean | 1 - Mathlib/CategoryTheory/Monoidal/Mon_.lean | 1 - Mathlib/CategoryTheory/Monoidal/NaturalTransformation.lean | 1 - Mathlib/CategoryTheory/MorphismProperty/IsInvertedBy.lean | 2 -- Mathlib/CategoryTheory/Pi/Basic.lean | 1 - Mathlib/CategoryTheory/Preadditive/Mat.lean | 3 --- Mathlib/CategoryTheory/Sites/Sheaf.lean | 1 - Mathlib/CategoryTheory/Subobject/Basic.lean | 3 --- Mathlib/CategoryTheory/Yoneda.lean | 4 ---- Mathlib/Geometry/RingedSpace/PresheafedSpace.lean | 1 - Mathlib/Geometry/RingedSpace/SheafedSpace.lean | 1 - Mathlib/RepresentationTheory/Action/Basic.lean | 1 - 33 files changed, 61 deletions(-) diff --git a/Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean b/Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean index c3f9a38ef087e..7d7e0e34c73af 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean @@ -365,7 +365,6 @@ def embeddingLiftIso (F : C ⥤ D) : embedding R C ⋙ lift R F ≅ F := /-- Two `R`-linear functors out of the `R`-linear completion are isomorphic iff their compositions with the embedding functor are isomorphic. -/ --- Porting note (#11182): used to be @[ext] def ext {F G : Free R C ⥤ D} [F.Additive] [F.Linear R] [G.Additive] [G.Linear R] (α : embedding R C ⋙ F ≅ embedding R C ⋙ G) : F ≅ G := NatIso.ofComponents (fun X => α.app X) diff --git a/Mathlib/Algebra/Group/AddChar.lean b/Mathlib/Algebra/Group/AddChar.lean index f9a1d383c185c..23df07dbc7f89 100644 --- a/Mathlib/Algebra/Group/AddChar.lean +++ b/Mathlib/Algebra/Group/AddChar.lean @@ -88,7 +88,6 @@ instance instFunLike : FunLike (AddChar A M) A M where coe := AddChar.toFun coe_injective' φ ψ h := by cases φ; cases ψ; congr --- Porting note (#5229): added. @[ext] lemma ext (f g : AddChar A M) (h : ∀ x : A, f x = g x) : f = g := DFunLike.ext f g h diff --git a/Mathlib/Algebra/Homology/HomologicalComplex.lean b/Mathlib/Algebra/Homology/HomologicalComplex.lean index 6ea13712672bd..441c75acfbc02 100644 --- a/Mathlib/Algebra/Homology/HomologicalComplex.lean +++ b/Mathlib/Algebra/Homology/HomologicalComplex.lean @@ -244,7 +244,6 @@ instance : Category (HomologicalComplex V c) where end --- Porting note (#5229): added because `Hom.ext` is not triggered automatically @[ext] lemma hom_ext {C D : HomologicalComplex V c} (f g : C ⟶ D) (h : ∀ i, f.f i = g.f i) : f = g := by diff --git a/Mathlib/Algebra/Polynomial/Laurent.lean b/Mathlib/Algebra/Polynomial/Laurent.lean index 949d81ecba031..c231267f42556 100644 --- a/Mathlib/Algebra/Polynomial/Laurent.lean +++ b/Mathlib/Algebra/Polynomial/Laurent.lean @@ -87,7 +87,6 @@ scoped[LaurentPolynomial] notation:9000 R "[T;T⁻¹]" => LaurentPolynomial R open LaurentPolynomial --- Porting note (#5229): `ext` no longer applies `Finsupp.ext` automatically @[ext] theorem LaurentPolynomial.ext [Semiring R] {p q : R[T;T⁻¹]} (h : ∀ a, p a = q a) : p = q := Finsupp.ext h diff --git a/Mathlib/AlgebraicTopology/SimplexCategory.lean b/Mathlib/AlgebraicTopology/SimplexCategory.lean index ac2101ce20054..72bc3cc192913 100644 --- a/Mathlib/AlgebraicTopology/SimplexCategory.lean +++ b/Mathlib/AlgebraicTopology/SimplexCategory.lean @@ -142,7 +142,6 @@ lemma id_toOrderHom (a : SimplexCategory) : lemma comp_toOrderHom {a b c : SimplexCategory} (f : a ⟶ b) (g : b ⟶ c) : (f ≫ g).toOrderHom = g.toOrderHom.comp f.toOrderHom := rfl --- Porting note (#5229): added because `Hom.ext'` is not triggered automatically @[ext] theorem Hom.ext {a b : SimplexCategory} (f g : a ⟶ b) : f.toOrderHom = g.toOrderHom → f = g := diff --git a/Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean b/Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean index b11a417a297e7..32b22e7b91ce5 100644 --- a/Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean +++ b/Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean @@ -57,7 +57,6 @@ instance hasColimits : HasColimits SSet := by dsimp only [SSet] infer_instance --- Porting note (#5229): added an `ext` lemma. @[ext] lemma hom_ext {X Y : SSet} {f g : X ⟶ Y} (w : ∀ n, f.app n = g.app n) : f = g := SimplicialObject.hom_ext _ _ w @@ -354,7 +353,6 @@ simplicial sets. -/ def Truncated.uliftFunctor (k : ℕ) : SSet.Truncated.{u} k ⥤ SSet.Truncated.{max u v} k := (whiskeringRight _ _ _).obj CategoryTheory.uliftFunctor.{v, u} --- Porting note (#5229): added an `ext` lemma. @[ext] lemma Truncated.hom_ext {n : ℕ} {X Y : Truncated n} {f g : X ⟶ Y} (w : ∀ n, f.app n = g.app n) : f = g := diff --git a/Mathlib/AlgebraicTopology/SplitSimplicialObject.lean b/Mathlib/AlgebraicTopology/SplitSimplicialObject.lean index 48c6023dd1053..7d8574e79627f 100644 --- a/Mathlib/AlgebraicTopology/SplitSimplicialObject.lean +++ b/Mathlib/AlgebraicTopology/SplitSimplicialObject.lean @@ -351,7 +351,6 @@ variable {C} namespace Split --- Porting note (#5229): added as `Hom.ext` is not triggered automatically @[ext] theorem hom_ext {S₁ S₂ : Split C} (Φ₁ Φ₂ : S₁ ⟶ S₂) (h : ∀ n : ℕ, Φ₁.f n = Φ₂.f n) : Φ₁ = Φ₂ := Hom.ext _ _ h diff --git a/Mathlib/CategoryTheory/Abelian/Pseudoelements.lean b/Mathlib/CategoryTheory/Abelian/Pseudoelements.lean index 7c2ed53010962..7bb3ceaced8e1 100644 --- a/Mathlib/CategoryTheory/Abelian/Pseudoelements.lean +++ b/Mathlib/CategoryTheory/Abelian/Pseudoelements.lean @@ -258,12 +258,6 @@ theorem zero_morphism_ext {P Q : C} (f : P ⟶ Q) : (∀ a, f a = 0) → f = 0 : theorem zero_morphism_ext' {P Q : C} (f : P ⟶ Q) : (∀ a, f a = 0) → 0 = f := Eq.symm ∘ zero_morphism_ext f --- Porting note (#11182): these are no longer valid as `ext` lemmas. --- scoped[Pseudoelement] --- attribute [ext] --- CategoryTheory.Abelian.Pseudoelement.zero_morphism_ext --- CategoryTheory.Abelian.Pseudoelement.zero_morphism_ext' - theorem eq_zero_iff {P Q : C} (f : P ⟶ Q) : f = 0 ↔ ∀ a, f a = 0 := ⟨fun h a => by simp [h], zero_morphism_ext _⟩ diff --git a/Mathlib/CategoryTheory/Bicategory/NaturalTransformation/Oplax.lean b/Mathlib/CategoryTheory/Bicategory/NaturalTransformation/Oplax.lean index 2d6217112669a..ee6b3cb99b45e 100644 --- a/Mathlib/CategoryTheory/Bicategory/NaturalTransformation/Oplax.lean +++ b/Mathlib/CategoryTheory/Bicategory/NaturalTransformation/Oplax.lean @@ -245,7 +245,6 @@ instance category (F G : OplaxFunctor B C) : Category (F ⟶ G) where id := Modification.id comp := Modification.vcomp --- Porting note (#5229): duplicating the `ext` lemma. @[ext] lemma ext {F G : OplaxFunctor B C} {α β : F ⟶ G} {m n : α ⟶ β} (w : ∀ b, m.app b = n.app b) : m = n := by diff --git a/Mathlib/CategoryTheory/Comma/Basic.lean b/Mathlib/CategoryTheory/Comma/Basic.lean index 74eebe08f5cf9..192ef0e4606e0 100644 --- a/Mathlib/CategoryTheory/Comma/Basic.lean +++ b/Mathlib/CategoryTheory/Comma/Basic.lean @@ -106,8 +106,6 @@ section variable {X Y Z : Comma L R} {f : X ⟶ Y} {g : Y ⟶ Z} --- Porting note (#5229): this lemma was added because `CommaMorphism.ext` --- was not triggered automatically @[ext] lemma hom_ext (f g : X ⟶ Y) (h₁ : f.left = g.left) (h₂ : f.right = g.right) : f = g := CommaMorphism.ext h₁ h₂ diff --git a/Mathlib/CategoryTheory/Comma/StructuredArrow.lean b/Mathlib/CategoryTheory/Comma/StructuredArrow.lean index 3ee2b719536a8..d86b41669eacf 100644 --- a/Mathlib/CategoryTheory/Comma/StructuredArrow.lean +++ b/Mathlib/CategoryTheory/Comma/StructuredArrow.lean @@ -51,8 +51,6 @@ def proj (S : D) (T : C ⥤ D) : StructuredArrow S T ⥤ C := variable {S S' S'' : D} {Y Y' Y'' : C} {T T' : C ⥤ D} --- Porting note (#5229): this lemma was added because `Comma.hom_ext` --- was not triggered automatically @[ext] lemma hom_ext {X Y : StructuredArrow S T} (f g : X ⟶ Y) (h : f.right = g.right) : f = g := CommaMorphism.ext (Subsingleton.elim _ _) h @@ -400,8 +398,6 @@ def proj (S : C ⥤ D) (T : D) : CostructuredArrow S T ⥤ C := variable {T T' T'' : D} {Y Y' Y'' : C} {S S' : C ⥤ D} --- Porting note (#5229): this lemma was added because `Comma.hom_ext` --- was not triggered automatically @[ext] lemma hom_ext {X Y : CostructuredArrow S T} (f g : X ⟶ Y) (h : f.left = g.left) : f = g := CommaMorphism.ext h (Subsingleton.elim _ _) diff --git a/Mathlib/CategoryTheory/Elements.lean b/Mathlib/CategoryTheory/Elements.lean index 588561a7941f8..cf261e1d8d7fc 100644 --- a/Mathlib/CategoryTheory/Elements.lean +++ b/Mathlib/CategoryTheory/Elements.lean @@ -47,7 +47,6 @@ def Functor.Elements (F : C ⥤ Type w) := /-- Constructor for the type `F.Elements` when `F` is a functor to types. -/ abbrev Functor.elementsMk (F : C ⥤ Type w) (X : C) (x : F.obj X) : F.Elements := ⟨X, x⟩ --- Porting note (#5229): added because Sigma.ext would be triggered automatically lemma Functor.Elements.ext {F : C ⥤ Type w} (x y : F.Elements) (h₁ : x.fst = y.fst) (h₂ : F.map (eqToHom h₁) x.snd = y.snd) : x = y := by cases x diff --git a/Mathlib/CategoryTheory/Endomorphism.lean b/Mathlib/CategoryTheory/Endomorphism.lean index 4845f0a202a03..0ab626814ae99 100644 --- a/Mathlib/CategoryTheory/Endomorphism.lean +++ b/Mathlib/CategoryTheory/Endomorphism.lean @@ -114,7 +114,6 @@ def Aut (X : C) := X ≅ X namespace Aut --- Porting note (#5229): added because `Iso.ext` is not triggered automatically @[ext] lemma ext {X : C} {φ₁ φ₂ : Aut X} (h : φ₁.hom = φ₂.hom) : φ₁ = φ₂ := Iso.ext h diff --git a/Mathlib/CategoryTheory/Functor/Category.lean b/Mathlib/CategoryTheory/Functor/Category.lean index e1bd5dc1cfab1..7a1687ae5667d 100644 --- a/Mathlib/CategoryTheory/Functor/Category.lean +++ b/Mathlib/CategoryTheory/Functor/Category.lean @@ -48,10 +48,6 @@ instance Functor.category : Category.{max u₁ v₂} (C ⥤ D) where namespace NatTrans --- Porting note (#5229): the behaviour of `ext` has changed here. --- We need to provide a copy of the `NatTrans.ext` lemma, --- written in terms of `F ⟶ G` rather than `NatTrans F G`, --- or `ext` will not retrieve it from the cache. @[ext] theorem ext' {α β : F ⟶ G} (w : α.app = β.app) : α = β := NatTrans.ext w diff --git a/Mathlib/CategoryTheory/Idempotents/Karoubi.lean b/Mathlib/CategoryTheory/Idempotents/Karoubi.lean index abcded134084e..b08afa205a73a 100644 --- a/Mathlib/CategoryTheory/Idempotents/Karoubi.lean +++ b/Mathlib/CategoryTheory/Idempotents/Karoubi.lean @@ -101,7 +101,6 @@ theorem hom_ext_iff {P Q : Karoubi C} {f g : P ⟶ Q} : f = g ↔ f.f = g.f := b rw [h] · apply Hom.ext --- Porting note (#5229): added because `Hom.ext` is not triggered automatically @[ext] theorem hom_ext {P Q : Karoubi C} (f g : P ⟶ Q) (h : f.f = g.f) : f = g := by simpa [hom_ext_iff] using h diff --git a/Mathlib/CategoryTheory/Iso.lean b/Mathlib/CategoryTheory/Iso.lean index 7ce50b850ef49..6e0dd598bc818 100644 --- a/Mathlib/CategoryTheory/Iso.lean +++ b/Mathlib/CategoryTheory/Iso.lean @@ -305,7 +305,6 @@ instance (priority := 100) mono_of_iso (f : X ⟶ Y) [IsIso f] : Mono f where rw [← Category.comp_id g, ← Category.comp_id h, ← IsIso.hom_inv_id f, ← Category.assoc, w, ← Category.assoc] --- Porting note (#11182): `@[ext]` used to accept lemmas like this. Now we add an aesop rule @[aesop apply safe (rule_sets := [CategoryTheory])] theorem inv_eq_of_hom_inv_id {f : X ⟶ Y} [IsIso f] {g : Y ⟶ X} (hom_inv_id : f ≫ g = 𝟙 X) : inv f = g := by @@ -317,7 +316,6 @@ theorem inv_eq_of_inv_hom_id {f : X ⟶ Y} [IsIso f] {g : Y ⟶ X} (inv_hom_id : apply (cancel_mono f).mp simp [inv_hom_id] --- Porting note (#11182): `@[ext]` used to accept lemmas like this. @[aesop apply safe (rule_sets := [CategoryTheory])] theorem eq_inv_of_hom_inv_id {f : X ⟶ Y} [IsIso f] {g : Y ⟶ X} (hom_inv_id : f ≫ g = 𝟙 X) : g = inv f := @@ -446,12 +444,10 @@ theorem isIso_of_comp_hom_eq_id (g : X ⟶ Y) [IsIso g] {f : Y ⟶ X} (h : f ≫ namespace Iso --- Porting note (#11182): `@[ext]` used to accept lemmas like this. @[aesop apply safe (rule_sets := [CategoryTheory])] theorem inv_ext {f : X ≅ Y} {g : Y ⟶ X} (hom_inv_id : f.hom ≫ g = 𝟙 X) : f.inv = g := ((hom_comp_eq_id f).1 hom_inv_id).symm --- Porting note (#11182): `@[ext]` used to accept lemmas like this. @[aesop apply safe (rule_sets := [CategoryTheory])] theorem inv_ext' {f : X ≅ Y} {g : Y ⟶ X} (hom_inv_id : f.hom ≫ g = 𝟙 X) : g = f.inv := (hom_comp_eq_id f).1 hom_inv_id diff --git a/Mathlib/CategoryTheory/Limits/Cones.lean b/Mathlib/CategoryTheory/Limits/Cones.lean index 098f9f7adcb6a..36d559c2f7911 100644 --- a/Mathlib/CategoryTheory/Limits/Cones.lean +++ b/Mathlib/CategoryTheory/Limits/Cones.lean @@ -283,7 +283,6 @@ namespace Cones /-- To give an isomorphism between cones, it suffices to give an isomorphism between their vertices which commutes with the cone maps. -/ --- Porting note (#11182): `@[ext]` used to accept lemmas like this. Now we add an aesop rule @[aesop apply safe (rule_sets := [CategoryTheory]), simps] def ext {c c' : Cone F} (φ : c.pt ≅ c'.pt) (w : ∀ j, c.π.app j = φ.hom ≫ c'.π.app j := by aesop_cat) : c ≅ c' where @@ -484,7 +483,6 @@ namespace Cocones /-- To give an isomorphism between cocones, it suffices to give an isomorphism between their vertices which commutes with the cocone maps. -/ --- Porting note (#11182): `@[ext]` used to accept lemmas like this. Now we add an aesop rule @[aesop apply safe (rule_sets := [CategoryTheory]), simps] def ext {c c' : Cocone F} (φ : c.pt ≅ c'.pt) (w : ∀ j, c.ι.app j ≫ φ.hom = c'.ι.app j := by aesop_cat) : c ≅ c' where diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean b/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean index 2d16499600981..f290e6be94303 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean @@ -122,7 +122,6 @@ namespace Bicones /-- To give an isomorphism between cocones, it suffices to give an isomorphism between their vertices which commutes with the cocone maps. -/ --- Porting note (#11182): `@[ext]` used to accept lemmas like this. Now we add an aesop rule @[aesop apply safe (rule_sets := [CategoryTheory]), simps] def ext {c c' : Bicone F} (φ : c.pt ≅ c'.pt) (wι : ∀ j, c.ι j ≫ φ.hom = c'.ι j := by aesop_cat) @@ -1119,7 +1118,6 @@ namespace BinaryBicones /-- To give an isomorphism between cocones, it suffices to give an isomorphism between their vertices which commutes with the cocone maps. -/ --- Porting note (#11182): `@[ext]` used to accept lemmas like this. Now we add an aesop rule @[aesop apply safe (rule_sets := [CategoryTheory]), simps] def ext {P Q : C} {c c' : BinaryBicone P Q} (φ : c.pt ≅ c'.pt) (winl : c.inl ≫ φ.hom = c'.inl := by aesop_cat) diff --git a/Mathlib/CategoryTheory/Monoidal/Bimod.lean b/Mathlib/CategoryTheory/Monoidal/Bimod.lean index b6d566f92f32a..6182c9de23c17 100644 --- a/Mathlib/CategoryTheory/Monoidal/Bimod.lean +++ b/Mathlib/CategoryTheory/Monoidal/Bimod.lean @@ -120,7 +120,6 @@ instance : Category (Bimod A B) where id := id' comp f g := comp f g --- Porting note (#5229): added because `Hom.ext` is not triggered automatically @[ext] lemma hom_ext {M N : Bimod A B} (f g : M ⟶ N) (h : f.hom = g.hom) : f = g := Hom.ext h diff --git a/Mathlib/CategoryTheory/Monoidal/Braided/Basic.lean b/Mathlib/CategoryTheory/Monoidal/Braided/Basic.lean index c0dfcace9f9a6..ee91313dea7f3 100644 --- a/Mathlib/CategoryTheory/Monoidal/Braided/Basic.lean +++ b/Mathlib/CategoryTheory/Monoidal/Braided/Basic.lean @@ -399,7 +399,6 @@ def comp (F : LaxBraidedFunctor C D) (G : LaxBraidedFunctor D E) : LaxBraidedFun instance categoryLaxBraidedFunctor : Category (LaxBraidedFunctor C D) := InducedCategory.category LaxBraidedFunctor.toLaxMonoidalFunctor --- Porting note (#5229): added, as `MonoidalNatTrans.ext` does not apply to morphisms. @[ext] lemma ext' {F G : LaxBraidedFunctor C D} {α β : F ⟶ G} (w : ∀ X : C, α.app X = β.app X) : α = β := MonoidalNatTrans.ext (funext w) @@ -463,7 +462,6 @@ def comp (F : BraidedFunctor C D) (G : BraidedFunctor D E) : BraidedFunctor C E instance categoryBraidedFunctor : Category (BraidedFunctor C D) := InducedCategory.category BraidedFunctor.toMonoidalFunctor --- Porting note (#5229): added, as `MonoidalNatTrans.ext` does not apply to morphisms. @[ext] lemma ext' {F G : BraidedFunctor C D} {α β : F ⟶ G} (w : ∀ X : C, α.app X = β.app X) : α = β := MonoidalNatTrans.ext (funext w) diff --git a/Mathlib/CategoryTheory/Monoidal/CommMon_.lean b/Mathlib/CategoryTheory/Monoidal/CommMon_.lean index b8a3c549ea911..95ff7663a6476 100644 --- a/Mathlib/CategoryTheory/Monoidal/CommMon_.lean +++ b/Mathlib/CategoryTheory/Monoidal/CommMon_.lean @@ -50,8 +50,6 @@ theorem comp_hom {R S T : CommMon_ C} (f : R ⟶ S) (g : S ⟶ T) : Mon_.Hom.hom (f ≫ g) = f.hom ≫ g.hom := rfl --- Porting note (#5229): added because `Mon_.Hom.ext` is not triggered automatically --- for morphisms in `CommMon_ C` @[ext] lemma hom_ext {A B : CommMon_ C} (f g : A ⟶ B) (h : f.hom = g.hom) : f = g := Mon_.Hom.ext h diff --git a/Mathlib/CategoryTheory/Monoidal/Mod_.lean b/Mathlib/CategoryTheory/Monoidal/Mod_.lean index bee5d1a324c2f..ec6327ca22917 100644 --- a/Mathlib/CategoryTheory/Monoidal/Mod_.lean +++ b/Mathlib/CategoryTheory/Monoidal/Mod_.lean @@ -57,7 +57,6 @@ instance : Category (Mod_ A) where id := id comp f g := comp f g --- Porting note (#5229): added because `Hom.ext` is not triggered automatically @[ext] lemma hom_ext {M N : Mod_ A} (f₁ f₂ : M ⟶ N) (h : f₁.hom = f₂.hom) : f₁ = f₂ := Hom.ext h diff --git a/Mathlib/CategoryTheory/Monoidal/Mon_.lean b/Mathlib/CategoryTheory/Monoidal/Mon_.lean index 82f033d04a045..fc0cec7c2faf7 100644 --- a/Mathlib/CategoryTheory/Monoidal/Mon_.lean +++ b/Mathlib/CategoryTheory/Monoidal/Mon_.lean @@ -104,7 +104,6 @@ instance : Category (Mon_ C) where id := id comp f g := comp f g --- Porting note (#5229): added, as `Hom.ext` does not apply to a morphism. @[ext] lemma ext {X Y : Mon_ C} {f g : X ⟶ Y} (w : f.hom = g.hom) : f = g := Hom.ext w diff --git a/Mathlib/CategoryTheory/Monoidal/NaturalTransformation.lean b/Mathlib/CategoryTheory/Monoidal/NaturalTransformation.lean index 4d73ab1b96ac5..d41970b459890 100644 --- a/Mathlib/CategoryTheory/Monoidal/NaturalTransformation.lean +++ b/Mathlib/CategoryTheory/Monoidal/NaturalTransformation.lean @@ -81,7 +81,6 @@ theorem comp_toNatTrans_lax {F G H : LaxMonoidalFunctor C D} {α : F ⟶ G} {β instance categoryMonoidalFunctor : Category (MonoidalFunctor C D) := InducedCategory.category MonoidalFunctor.toLaxMonoidalFunctor --- Porting note (#5229): added, as `MonoidalNatTrans.ext` does not apply to morphisms. @[ext] lemma ext' {F G : LaxMonoidalFunctor C D} {α β : F ⟶ G} (w : ∀ X : C, α.app X = β.app X) : α = β := MonoidalNatTrans.ext (funext w) diff --git a/Mathlib/CategoryTheory/MorphismProperty/IsInvertedBy.lean b/Mathlib/CategoryTheory/MorphismProperty/IsInvertedBy.lean index e4d78ec3864b8..a20da408bc49c 100644 --- a/Mathlib/CategoryTheory/MorphismProperty/IsInvertedBy.lean +++ b/Mathlib/CategoryTheory/MorphismProperty/IsInvertedBy.lean @@ -104,8 +104,6 @@ lemma FunctorsInverting.ext {W : MorphismProperty C} {F₁ F₂ : FunctorsInvert instance (W : MorphismProperty C) (D : Type*) [Category D] : Category (FunctorsInverting W D) := FullSubcategory.category _ --- Porting note (#5229): add another `@[ext]` lemma --- since `ext` can't see through the definition to use `NatTrans.ext`. @[ext] lemma FunctorsInverting.hom_ext {W : MorphismProperty C} {F₁ F₂ : FunctorsInverting W D} {α β : F₁ ⟶ F₂} (h : α.app = β.app) : α = β := diff --git a/Mathlib/CategoryTheory/Pi/Basic.lean b/Mathlib/CategoryTheory/Pi/Basic.lean index 8515247f7d5ba..477e2fe2a21e5 100644 --- a/Mathlib/CategoryTheory/Pi/Basic.lean +++ b/Mathlib/CategoryTheory/Pi/Basic.lean @@ -49,7 +49,6 @@ theorem comp_apply {X Y Z : ∀ i, C i} (f : X ⟶ Y) (g : Y ⟶ Z) (i) : (f ≫ g : ∀ i, X i ⟶ Z i) i = f i ≫ g i := rfl --- Porting note (#5229): need to add an additional `ext` lemma. @[ext] lemma ext {X Y : ∀ i, C i} {f g : X ⟶ Y} (w : ∀ i, f i = g i) : f = g := funext (w ·) diff --git a/Mathlib/CategoryTheory/Preadditive/Mat.lean b/Mathlib/CategoryTheory/Preadditive/Mat.lean index 9d44442e07dd3..8ab0b6f6a6063 100644 --- a/Mathlib/CategoryTheory/Preadditive/Mat.lean +++ b/Mathlib/CategoryTheory/Preadditive/Mat.lean @@ -105,7 +105,6 @@ instance : Category.{v₁} (Mat_ C) where simp_rw [Hom.comp, sum_comp, comp_sum, Category.assoc] rw [Finset.sum_comm] --- Porting note (#5229): added because `DMatrix.ext` is not triggered automatically @[ext] theorem hom_ext {M N : Mat_ C} (f g : M ⟶ N) (H : ∀ i j, f i j = g i j) : f = g := DMatrix.ext_iff.mp H @@ -432,7 +431,6 @@ def liftUnique (F : C ⥤ D) [Functor.Additive F] (L : Mat_ C ⥤ D) [Functor.Ad dsimp simpa using α.hom.naturality (f j k) --- Porting note (#11182): removed @[ext] as the statement is not an equality -- TODO is there some uniqueness statement for the natural isomorphism in `liftUnique`? /-- Two additive functors `Mat_ C ⥤ D` are naturally isomorphic if their precompositions with `embedding C` are naturally isomorphic as functors `C ⥤ D`. -/ @@ -505,7 +503,6 @@ section variable {R : Type u} [Semiring R] --- Porting note (#5229): added because `Matrix.ext` is not triggered automatically @[ext] theorem hom_ext {X Y : Mat R} (f g : X ⟶ Y) (h : ∀ i j, f i j = g i j) : f = g := Matrix.ext_iff.mp h diff --git a/Mathlib/CategoryTheory/Sites/Sheaf.lean b/Mathlib/CategoryTheory/Sites/Sheaf.lean index 99d958c7ef686..10da297aaf6d1 100644 --- a/Mathlib/CategoryTheory/Sites/Sheaf.lean +++ b/Mathlib/CategoryTheory/Sites/Sheaf.lean @@ -332,7 +332,6 @@ instance instCategorySheaf : Category (Sheaf J A) where instance (X : Sheaf J A) : Inhabited (Hom X X) := ⟨𝟙 X⟩ --- Porting note (#5229): added because `Sheaf.Hom.ext` was not triggered automatically @[ext] lemma hom_ext {X Y : Sheaf J A} (x y : X ⟶ Y) (h : x.val = y.val) : x = y := Sheaf.Hom.ext h diff --git a/Mathlib/CategoryTheory/Subobject/Basic.lean b/Mathlib/CategoryTheory/Subobject/Basic.lean index 34c23433816f0..3afc05fb08654 100644 --- a/Mathlib/CategoryTheory/Subobject/Basic.lean +++ b/Mathlib/CategoryTheory/Subobject/Basic.lean @@ -255,21 +255,18 @@ theorem eq_of_comm {B : C} {X Y : Subobject B} (f : (X : C) ≅ (Y : C)) (w : f.hom ≫ Y.arrow = X.arrow) : X = Y := le_antisymm (le_of_comm f.hom w) <| le_of_comm f.inv <| f.inv_comp_eq.2 w.symm --- Porting note (#11182): removed @[ext] /-- To show that two subobjects are equal, it suffices to exhibit an isomorphism commuting with the arrows. -/ theorem eq_mk_of_comm {B A : C} {X : Subobject B} (f : A ⟶ B) [Mono f] (i : (X : C) ≅ A) (w : i.hom ≫ f = X.arrow) : X = mk f := eq_of_comm (i.trans (underlyingIso f).symm) <| by simp [w] --- Porting note (#11182): removed @[ext] /-- To show that two subobjects are equal, it suffices to exhibit an isomorphism commuting with the arrows. -/ theorem mk_eq_of_comm {B A : C} {X : Subobject B} (f : A ⟶ B) [Mono f] (i : A ≅ (X : C)) (w : i.hom ≫ X.arrow = f) : mk f = X := Eq.symm <| eq_mk_of_comm _ i.symm <| by rw [Iso.symm_hom, Iso.inv_comp_eq, w] --- Porting note (#11182): removed @[ext] /-- To show that two subobjects are equal, it suffices to exhibit an isomorphism commuting with the arrows. -/ theorem mk_eq_mk_of_comm {B A₁ A₂ : C} (f : A₁ ⟶ B) (g : A₂ ⟶ B) [Mono f] [Mono g] (i : A₁ ≅ A₂) diff --git a/Mathlib/CategoryTheory/Yoneda.lean b/Mathlib/CategoryTheory/Yoneda.lean index d056292dbff45..61b37dc22ad22 100644 --- a/Mathlib/CategoryTheory/Yoneda.lean +++ b/Mathlib/CategoryTheory/Yoneda.lean @@ -486,8 +486,6 @@ to `yoneda.op.obj X ⟶ F`, functorially in both `X` and `F`. def yonedaPairing : Cᵒᵖ × (Cᵒᵖ ⥤ Type v₁) ⥤ Type max u₁ v₁ := Functor.prod yoneda.op (𝟭 (Cᵒᵖ ⥤ Type v₁)) ⋙ Functor.hom (Cᵒᵖ ⥤ Type v₁) --- Porting note (#5229): we need to provide this `@[ext]` lemma separately, --- as `ext` will not look through the definition. @[ext] lemma yonedaPairingExt {X : Cᵒᵖ × (Cᵒᵖ ⥤ Type v₁)} {x y : (yonedaPairing C).obj X} (w : ∀ Y, x.app Y = y.app Y) : x = y := @@ -657,8 +655,6 @@ to `coyoneda.rightOp.obj X ⟶ F`, functorially in both `X` and `F`. def coyonedaPairing : C × (C ⥤ Type v₁) ⥤ Type max u₁ v₁ := Functor.prod coyoneda.rightOp (𝟭 (C ⥤ Type v₁)) ⋙ Functor.hom (C ⥤ Type v₁) --- Porting note (#5229): we need to provide this `@[ext]` lemma separately, --- as `ext` will not look through the definition. @[ext] lemma coyonedaPairingExt {X : C × (C ⥤ Type v₁)} {x y : (coyonedaPairing C).obj X} (w : ∀ Y, x.app Y = y.app Y) : x = y := diff --git a/Mathlib/Geometry/RingedSpace/PresheafedSpace.lean b/Mathlib/Geometry/RingedSpace/PresheafedSpace.lean index f4762b2bced5c..12c2d277fbd66 100644 --- a/Mathlib/Geometry/RingedSpace/PresheafedSpace.lean +++ b/Mathlib/Geometry/RingedSpace/PresheafedSpace.lean @@ -160,7 +160,6 @@ variable {C} /-- Cast `Hom X Y` as an arrow `X ⟶ Y` of presheaves. -/ abbrev Hom.toPshHom {X Y : PresheafedSpace C} (f : Hom X Y) : X ⟶ Y := f --- Porting note (#5229): adding an `ext` lemma. @[ext (iff := false)] theorem ext {X Y : PresheafedSpace C} (α β : X ⟶ Y) (w : α.base = β.base) (h : α.c ≫ whiskerRight (eqToHom (by rw [w])) _ = β.c) : α = β := diff --git a/Mathlib/Geometry/RingedSpace/SheafedSpace.lean b/Mathlib/Geometry/RingedSpace/SheafedSpace.lean index 7b65ebaffd7eb..3af5f2d7e91d0 100644 --- a/Mathlib/Geometry/RingedSpace/SheafedSpace.lean +++ b/Mathlib/Geometry/RingedSpace/SheafedSpace.lean @@ -80,7 +80,6 @@ instance : Category (SheafedSpace C) := show Category (InducedCategory (PresheafedSpace C) SheafedSpace.toPresheafedSpace) by infer_instance --- Porting note (#5229): adding an `ext` lemma. @[ext (iff := false)] theorem ext {X Y : SheafedSpace C} (α β : X ⟶ Y) (w : α.base = β.base) (h : α.c ≫ whiskerRight (eqToHom (by rw [w])) _ = β.c) : α = β := diff --git a/Mathlib/RepresentationTheory/Action/Basic.lean b/Mathlib/RepresentationTheory/Action/Basic.lean index f70ce1fb09e29..2f15639455bdf 100644 --- a/Mathlib/RepresentationTheory/Action/Basic.lean +++ b/Mathlib/RepresentationTheory/Action/Basic.lean @@ -111,7 +111,6 @@ instance : Category (Action V G) where id M := Hom.id M comp f g := Hom.comp f g --- Porting note (#5229): added because `Hom.ext` is not triggered automatically @[ext] lemma hom_ext {M N : Action V G} (φ₁ φ₂ : M ⟶ N) (h : φ₁.hom = φ₂.hom) : φ₁ = φ₂ := Hom.ext h From 3ffd5e037b602dd5939615c5556c0464cae46492 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Calle=20S=C3=B6nne?= Date: Thu, 17 Oct 2024 15:41:44 +0000 Subject: [PATCH 097/505] feat(MorphismProperty/Presheaf): Add properties of the `relative` morphism property (#16142) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR adds some basic properties about the "presheaf" morphism property. For example, that it is stable under base change. Co-authored-by: Joël Riou --- .../MorphismProperty/Representable.lean | 58 ++++++++++++++++++- 1 file changed, 55 insertions(+), 3 deletions(-) diff --git a/Mathlib/CategoryTheory/MorphismProperty/Representable.lean b/Mathlib/CategoryTheory/MorphismProperty/Representable.lean index e2ff92296f5c9..df96aa39de010 100644 --- a/Mathlib/CategoryTheory/MorphismProperty/Representable.lean +++ b/Mathlib/CategoryTheory/MorphismProperty/Representable.lean @@ -25,9 +25,10 @@ Classically, a morphism `f : F ⟶ G` of presheaves is said to be representable In this file, we define a notion of relative representability which works with respect to any functor, and not just `yoneda`. The fact that a morphism `f : F ⟶ G` between presheaves is -representable in the classical case will then be given by `F.relativelyRepresentable f`. -` +representable in the classical case will then be given by `yoneda.relativelyRepresentable f`. + ## Main definitions + Throughout this file, `F : C ⥤ D` is a functor between categories `C` and `D`. * `Functor.relativelyRepresentable`: A morphism `f : X ⟶ Y` in `D` is said to be relatively @@ -312,7 +313,6 @@ category `Cᵒᵖ ⥤ Type v` satisfies the morphism property `P.presheaf` iff: * The morphism is representable. * For any morphism `g : F.obj a ⟶ G`, the property `P` holds for any represented pullback of `f` by `g`. - This is implemented as a special case of the more general notion of `P.relative`, to the case when the functor `F` is `yoneda`. -/ abbrev presheaf : MorphismProperty (Cᵒᵖ ⥤ Type v₁) := P.relative yoneda @@ -372,6 +372,58 @@ lemma relative_map_iff [F.Faithful] [F.Full] [PreservesLimitsOfShape WalkingCosp P.relative F (F.map f) ↔ P f := ⟨fun hf ↦ of_relative_map hf, fun hf ↦ relative_map hP hf⟩ +/-- If `P' : MorphismProperty C` is satisfied whenever `P` is, then also `P'.relative` is +satisfied whenever `P.relative` is. -/ +lemma relative_monotone {P' : MorphismProperty C} (h : P ≤ P') : + P.relative F ≤ P'.relative F := fun _ _ _ hf ↦ + ⟨hf.rep, fun _ _ g fst snd BC ↦ h _ (hf.property g fst snd BC)⟩ + +section + +variable (P) + +lemma relative_stableUnderBaseChange : StableUnderBaseChange (P.relative F) := + fun _ _ _ _ _ _ _ _ hfBC hg ↦ + ⟨stableUnderBaseChange F hfBC hg.rep, + fun _ _ _ _ _ BC ↦ hg.property _ _ _ (IsPullback.paste_horiz BC hfBC)⟩ + +instance relative_isStableUnderComposition [F.Faithful] [F.Full] [P.IsStableUnderComposition] : + IsStableUnderComposition (P.relative F) where + comp_mem {F G H} f g hf hg := by + refine ⟨comp_mem _ _ _ hf.1 hg.1, fun Z X p fst snd h ↦ ?_⟩ + rw [← hg.1.lift_snd (fst ≫ f) snd (by simpa using h.w)] + refine comp_mem _ _ _ (hf.property (hg.1.fst p) fst _ + (IsPullback.of_bot ?_ ?_ (hg.1.isPullback p))) (hg.property_snd p) + · rw [← Functor.map_comp, lift_snd] + exact h + · symm + apply hg.1.lift_fst + +instance relative_respectsIso : RespectsIso (P.relative F) := + (relative_stableUnderBaseChange P).respectsIso + +instance relative_isMultiplicative [F.Faithful] [F.Full] [P.IsMultiplicative] [P.RespectsIso] : + IsMultiplicative (P.relative F) where + id_mem X := relative.of_exists + (fun Y g ↦ ⟨Y, g, 𝟙 Y, by simpa using IsPullback.of_id_snd, id_mem _ _⟩) + +end + +/-- Morphisms satisfying `(monomorphism C).presheaf` are in particular monomorphisms. -/ +lemma presheaf_monomorphisms_le_monomorphisms : + (monomorphisms C).presheaf ≤ monomorphisms _ := fun F G f hf ↦ by + suffices ∀ {X : C} {a b : yoneda.obj X ⟶ F}, a ≫ f = b ≫ f → a = b from + ⟨fun _ _ h ↦ hom_ext_yoneda (fun _ _ ↦ this (by simp only [assoc, h]))⟩ + intro X a b h + /- It suffices to show that the lifts of `a` and `b` to morphisms + `X ⟶ hf.rep.pullback g` are equal, where `g = a ≫ f = a ≫ f`. -/ + suffices hf.rep.lift (g := a ≫ f) a (𝟙 X) (by simp) = + hf.rep.lift b (𝟙 X) (by simp [← h]) by + simpa using yoneda.congr_map this =≫ (hf.rep.fst (a ≫ f)) + -- This follows from the fact that the induced maps `hf.rep.pullback g ⟶ X` are mono. + have : Mono (hf.rep.snd (a ≫ f)) := hf.property_snd (a ≫ f) + simp only [← cancel_mono (hf.rep.snd (a ≫ f)), lift_snd] + end MorphismProperty end CategoryTheory From 27616824e4ad7714586263171a565683e6a6d7c9 Mon Sep 17 00:00:00 2001 From: Dagur Asgeirsson Date: Thu, 17 Oct 2024 15:41:47 +0000 Subject: [PATCH 098/505] chore(Profinite): remove simp lemma `toProfinite_obj` (#17155) --- Mathlib/Condensed/Discrete/Colimit.lean | 41 ++++++++----------- .../Category/LightProfinite/AsLimit.lean | 3 +- .../Category/LightProfinite/Basic.lean | 2 +- .../Topology/Category/Profinite/Basic.lean | 2 +- 4 files changed, 21 insertions(+), 27 deletions(-) diff --git a/Mathlib/Condensed/Discrete/Colimit.lean b/Mathlib/Condensed/Discrete/Colimit.lean index 2e2de66d0a095..75f3c6ee7dfeb 100644 --- a/Mathlib/Condensed/Discrete/Colimit.lean +++ b/Mathlib/Condensed/Discrete/Colimit.lean @@ -48,10 +48,7 @@ noncomputable def isColimitLocallyConstantPresheaf (hc : IsLimit c) [∀ i, Epi (h : fi.comap (c.π.app i) = fj.comap (c.π.app j)) obtain ⟨k, ki, kj, _⟩ := IsCofilteredOrEmpty.cone_objs i j refine ⟨⟨k⟩, ki.op, kj.op, ?_⟩ - dsimp only [comp_obj, op_obj, functorToPresheaves_obj_obj, CompHausLike.coe_of, - Functor.comp_map, op_map, Quiver.Hom.unop_op, functorToPresheaves_obj_map] - -- Note: we might want to remove the `simps` attribute from `FintypeCat.toProfinite`; keeping - -- `toProfinite_obj` in the `dsimp` block above causes the following `ext` to fail. + dsimp ext x obtain ⟨x, hx⟩ := ((Profinite.epi_iff_surjective (c.π.app k)).mp inferInstance) x rw [← hx] @@ -237,7 +234,7 @@ def isoFinYoneda : toProfinite.op ⋙ F ≅ finYoneda F := NatIso.ofComponents (fun X ↦ isoFinYonedaComponents F (toProfinite.obj X.unop)) fun _ ↦ by simp only [comp_obj, op_obj, finYoneda_obj, Functor.comp_map, op_map] ext - simp only [toProfinite_obj, types_comp_apply, isoFinYonedaComponents_hom_apply, finYoneda_map, + simp only [types_comp_apply, isoFinYonedaComponents_hom_apply, finYoneda_map, op_obj, Function.comp_apply, ← FunctorToTypes.map_comp_apply] rfl @@ -264,14 +261,14 @@ lemma isoLocallyConstantOfIsColimit_inv (X : Profinite.{u}ᵒᵖ ⥤ Type (u+1)) apply colimit.hom_ext intro ⟨Y, _, g⟩ simp? [locallyConstantIsoFinYoneda, isoFinYoneda, counitApp] says - simp only [comp_obj, CostructuredArrow.proj_obj, op_obj, toProfinite_obj, - functorToPresheaves_obj_obj, CompHausLike.coe_of, isoFinYoneda, locallyConstantIsoFinYoneda, - finYoneda_obj, LocallyConstant.toFun_eq_coe, NatTrans.comp_app, pointwiseLeftKanExtension_obj, - lanPresheafExt_inv, Iso.trans_inv, Iso.symm_inv, whiskerLeft_comp, lanPresheafNatIso_hom_app, - Opposite.op_unop, colimit.map_desc, id_eq, Functor.comp_map, op_map, colimit.ι_desc, - Cocones.precompose_obj_pt, Profinite.Extend.cocone_pt, Cocones.precompose_obj_ι, - Category.assoc, const_obj_obj, whiskerLeft_app, NatIso.ofComponents_hom_app, - NatIso.ofComponents_inv_app, Profinite.Extend.cocone_ι_app, counitApp, colimit.ι_desc_assoc] + simp only [comp_obj, CostructuredArrow.proj_obj, op_obj, functorToPresheaves_obj_obj, + isoFinYoneda, locallyConstantIsoFinYoneda, finYoneda_obj, LocallyConstant.toFun_eq_coe, + NatTrans.comp_app, pointwiseLeftKanExtension_obj, lanPresheafExt_inv, Iso.trans_inv, + Iso.symm_inv, whiskerLeft_comp, lanPresheafNatIso_hom_app, Opposite.op_unop, colimit.map_desc, + id_eq, Functor.comp_map, op_map, colimit.ι_desc, Cocones.precompose_obj_pt, + Profinite.Extend.cocone_pt, Cocones.precompose_obj_ι, Category.assoc, const_obj_obj, + whiskerLeft_app, NatIso.ofComponents_hom_app, NatIso.ofComponents_inv_app, + Profinite.Extend.cocone_ι_app, counitApp, colimit.ι_desc_assoc] erw [(counitApp.{u, u+1} X).naturality] simp only [← Category.assoc] congr @@ -318,8 +315,7 @@ noncomputable def isColimitLocallyConstantPresheaf (hc : IsLimit c) [∀ i, Epi (h : fi.comap (c.π.app i) = fj.comap (c.π.app j)) obtain ⟨k, ki, kj, _⟩ := IsCofilteredOrEmpty.cone_objs i j refine ⟨⟨k⟩, ki.op, kj.op, ?_⟩ - dsimp only [comp_obj, op_obj, functorToPresheaves_obj_obj, CompHausLike.coe_of, - Functor.comp_map, op_map, Quiver.Hom.unop_op, functorToPresheaves_obj_map] + dsimp ext x obtain ⟨x, hx⟩ := ((LightProfinite.epi_iff_surjective (c.π.app k)).mp inferInstance) x rw [← hx] @@ -541,14 +537,13 @@ lemma isoLocallyConstantOfIsColimit_inv (X : LightProfinite.{u}ᵒᵖ ⥤ Type u intro ⟨Y, _, g⟩ simp? [locallyConstantIsoFinYoneda, isoFinYoneda, counitApp] says simp only [comp_obj, CostructuredArrow.proj_obj, op_obj, functorToPresheaves_obj_obj, - toLightProfinite_obj_toTop_α, isoFinYoneda, locallyConstantIsoFinYoneda, finYoneda_obj, - LocallyConstant.toFun_eq_coe, NatTrans.comp_app, pointwiseLeftKanExtension_obj, - lanPresheafExt_inv, Iso.trans_inv, Iso.symm_inv, whiskerLeft_comp, lanPresheafNatIso_hom_app, - Opposite.op_unop, colimit.map_desc, id_eq, Functor.comp_map, op_map, colimit.ι_desc, - Cocones.precompose_obj_pt, LightProfinite.Extend.cocone_pt, Cocones.precompose_obj_ι, - Category.assoc, const_obj_obj, whiskerLeft_app, NatIso.ofComponents_hom_app, - NatIso.ofComponents_inv_app, LightProfinite.Extend.cocone_ι_app, counitApp, - colimit.ι_desc_assoc] + isoFinYoneda, locallyConstantIsoFinYoneda, finYoneda_obj, LocallyConstant.toFun_eq_coe, + NatTrans.comp_app, pointwiseLeftKanExtension_obj, lanPresheafExt_inv, Iso.trans_inv, + Iso.symm_inv, whiskerLeft_comp, lanPresheafNatIso_hom_app, Opposite.op_unop, colimit.map_desc, + id_eq, Functor.comp_map, op_map, colimit.ι_desc, Cocones.precompose_obj_pt, + LightProfinite.Extend.cocone_pt, Cocones.precompose_obj_ι, Category.assoc, const_obj_obj, + whiskerLeft_app, NatIso.ofComponents_hom_app, NatIso.ofComponents_inv_app, + LightProfinite.Extend.cocone_ι_app, counitApp, colimit.ι_desc_assoc] erw [(counitApp.{u, u} X).naturality] simp only [← Category.assoc] congr diff --git a/Mathlib/Topology/Category/LightProfinite/AsLimit.lean b/Mathlib/Topology/Category/LightProfinite/AsLimit.lean index 48104258a6815..bb092cf45b600 100644 --- a/Mathlib/Topology/Category/LightProfinite/AsLimit.lean +++ b/Mathlib/Topology/Category/LightProfinite/AsLimit.lean @@ -79,8 +79,7 @@ abbrev proj (n : ℕ) : S ⟶ S.diagram.obj ⟨n⟩ := S.asLimitCone.π.app ⟨n lemma lightToProfinite_map_proj_eq (n : ℕ) : lightToProfinite.map (S.proj n) = (lightToProfinite.obj S).asLimitCone.π.app _ := by - simp? says simp only [toCompHausLike_obj, Functor.comp_obj, - FintypeCat.toLightProfinite_obj_toTop_α, toCompHausLike_map, coe_of] + simp only [toCompHausLike_obj, Functor.comp_obj, toCompHausLike_map, coe_of] let c : Cone (S.diagram ⋙ lightToProfinite) := S.toLightDiagram.cone let hc : IsLimit c := S.toLightDiagram.isLimit exact liftedLimitMapsToOriginal_inv_map_π hc _ diff --git a/Mathlib/Topology/Category/LightProfinite/Basic.lean b/Mathlib/Topology/Category/LightProfinite/Basic.lean index 73dbb2fa55ef6..7d300f3905088 100644 --- a/Mathlib/Topology/Category/LightProfinite/Basic.lean +++ b/Mathlib/Topology/Category/LightProfinite/Basic.lean @@ -100,7 +100,7 @@ attribute [local instance] FintypeCat.discreteTopology /-- The natural functor from `Fintype` to `LightProfinite`, endowing a finite type with the discrete topology. -/ -@[simps!] +@[simps! map_apply] def FintypeCat.toLightProfinite : FintypeCat ⥤ LightProfinite where obj A := LightProfinite.of A map f := ⟨f, by continuity⟩ diff --git a/Mathlib/Topology/Category/Profinite/Basic.lean b/Mathlib/Topology/Category/Profinite/Basic.lean index 1faa3653c9a74..45aaebd8b6c14 100644 --- a/Mathlib/Topology/Category/Profinite/Basic.lean +++ b/Mathlib/Topology/Category/Profinite/Basic.lean @@ -140,7 +140,7 @@ attribute [local instance] FintypeCat.discreteTopology /-- The natural functor from `Fintype` to `Profinite`, endowing a finite type with the discrete topology. -/ -@[simps] +@[simps map_apply] def FintypeCat.toProfinite : FintypeCat ⥤ Profinite where obj A := Profinite.of A map f := ⟨f, by continuity⟩ From d9aad45391e594b51dfd27fdf9c579e105a0183c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Thu, 17 Oct 2024 15:41:48 +0000 Subject: [PATCH 099/505] feat(CategoryTheory/Shift): shifted morphisms in the opposite category (#17546) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit When `C` is a triangulated category, we introduce a bijection `ShiftedHom.opEquiv` which identifies `X ⟶ Y⟦n⟧` and `op Y ⟶ (op X)⟦n⟧` for any objects `X` and `Y` in `C`, and `n : ℤ`. Along with the compatibilities that are obtained, this shall be the main ingredient in the construction of the (contravariant) long exact sequence of `Ext` in abelian categories #15092. --- Mathlib.lean | 1 + Mathlib/CategoryTheory/Shift/Basic.lean | 65 +++++++- .../Shift/ShiftedHomOpposite.lean | 155 ++++++++++++++++++ .../CategoryTheory/Triangulated/Opposite.lean | 55 ++++++- .../CategoryTheory/Triangulated/Yoneda.lean | 40 ++++- Mathlib/Combinatorics/Quiver/Basic.lean | 9 + 6 files changed, 320 insertions(+), 5 deletions(-) create mode 100644 Mathlib/CategoryTheory/Shift/ShiftedHomOpposite.lean diff --git a/Mathlib.lean b/Mathlib.lean index 155e528def04f..b43069a00a054 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1858,6 +1858,7 @@ import Mathlib.CategoryTheory.Shift.Pullback import Mathlib.CategoryTheory.Shift.Quotient import Mathlib.CategoryTheory.Shift.ShiftSequence import Mathlib.CategoryTheory.Shift.ShiftedHom +import Mathlib.CategoryTheory.Shift.ShiftedHomOpposite import Mathlib.CategoryTheory.Shift.SingleFunctors import Mathlib.CategoryTheory.Sigma.Basic import Mathlib.CategoryTheory.Simple diff --git a/Mathlib/CategoryTheory/Shift/Basic.lean b/Mathlib/CategoryTheory/Shift/Basic.lean index c8b2f4ddfaec9..7acb86cbee29d 100644 --- a/Mathlib/CategoryTheory/Shift/Basic.lean +++ b/Mathlib/CategoryTheory/Shift/Basic.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2020 Kim Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Kim Morrison, Johan Commelin, Andrew Yang +Authors: Kim Morrison, Johan Commelin, Andrew Yang, Joël Riou -/ import Mathlib.Algebra.Group.Basic import Mathlib.CategoryTheory.Limits.Preserves.Shapes.Zero @@ -487,6 +487,69 @@ theorem shift_shiftFunctorCompIsoId_neg_add_cancel_inv_app (n : A) (X : C) : end +section + +variable (A) + +lemma shiftFunctorCompIsoId_zero_zero_hom_app (X : C) : + (shiftFunctorCompIsoId C 0 0 (add_zero 0)).hom.app X = + ((shiftFunctorZero C A).hom.app X)⟦0⟧' ≫ (shiftFunctorZero C A).hom.app X := by + simp [shiftFunctorCompIsoId, shiftFunctorAdd'_zero_add_inv_app] + +lemma shiftFunctorCompIsoId_zero_zero_inv_app (X : C) : + (shiftFunctorCompIsoId C 0 0 (add_zero 0)).inv.app X = + (shiftFunctorZero C A).inv.app X ≫ ((shiftFunctorZero C A).inv.app X)⟦0⟧' := by + simp [shiftFunctorCompIsoId, shiftFunctorAdd'_zero_add_hom_app] + +end + +section + +variable (m n p m' n' p' : A) (hm : m' + m = 0) (hn : n' + n = 0) (hp : p' + p = 0) + (h : m + n = p) + +lemma shiftFunctorCompIsoId_add'_inv_app : + (shiftFunctorCompIsoId C p' p hp).inv.app X = + (shiftFunctorCompIsoId C n' n hn).inv.app X ≫ + (shiftFunctorCompIsoId C m' m hm).inv.app (X⟦n'⟧)⟦n⟧' ≫ + (shiftFunctorAdd' C m n p h).inv.app (X⟦n'⟧⟦m'⟧) ≫ + ((shiftFunctorAdd' C n' m' p' + (by rw [← add_left_inj p, hp, ← h, add_assoc, + ← add_assoc m', hm, zero_add, hn])).inv.app X)⟦p⟧' := by + dsimp [shiftFunctorCompIsoId] + simp only [Functor.map_comp, Category.assoc] + congr 1 + erw [← NatTrans.naturality] + dsimp + rw [← cancel_mono ((shiftFunctorAdd' C p' p 0 hp).inv.app X), Iso.hom_inv_id_app, + Category.assoc, Category.assoc, Category.assoc, Category.assoc, + ← shiftFunctorAdd'_assoc_inv_app p' m n n' p 0 + (by rw [← add_left_inj n, hn, add_assoc, h, hp]) h (by rw [add_assoc, h, hp]), + ← Functor.map_comp_assoc, ← Functor.map_comp_assoc, ← Functor.map_comp_assoc, + Category.assoc, Category.assoc, + shiftFunctorAdd'_assoc_inv_app n' m' m p' 0 n' _ _ + (by rw [add_assoc, hm, add_zero]), Iso.hom_inv_id_app_assoc, + ← shiftFunctorAdd'_add_zero_hom_app, Iso.hom_inv_id_app, + Functor.map_id, Category.id_comp, Iso.hom_inv_id_app] + +lemma shiftFunctorCompIsoId_add'_hom_app : + (shiftFunctorCompIsoId C p' p hp).hom.app X = + ((shiftFunctorAdd' C n' m' p' + (by rw [← add_left_inj p, hp, ← h, add_assoc, + ← add_assoc m', hm, zero_add, hn])).hom.app X)⟦p⟧' ≫ + (shiftFunctorAdd' C m n p h).hom.app (X⟦n'⟧⟦m'⟧) ≫ + (shiftFunctorCompIsoId C m' m hm).hom.app (X⟦n'⟧)⟦n⟧' ≫ + (shiftFunctorCompIsoId C n' n hn).hom.app X := by + rw [← cancel_mono ((shiftFunctorCompIsoId C p' p hp).inv.app X), Iso.hom_inv_id_app, + shiftFunctorCompIsoId_add'_inv_app m n p m' n' p' hm hn hp h, + Category.assoc, Category.assoc, Category.assoc, Iso.hom_inv_id_app_assoc, + ← Functor.map_comp_assoc, Iso.hom_inv_id_app] + dsimp + rw [Functor.map_id, Category.id_comp, Iso.hom_inv_id_app_assoc, + ← Functor.map_comp, Iso.hom_inv_id_app, Functor.map_id] + +end + open CategoryTheory.Limits variable [HasZeroMorphisms C] diff --git a/Mathlib/CategoryTheory/Shift/ShiftedHomOpposite.lean b/Mathlib/CategoryTheory/Shift/ShiftedHomOpposite.lean new file mode 100644 index 0000000000000..5f21658547de3 --- /dev/null +++ b/Mathlib/CategoryTheory/Shift/ShiftedHomOpposite.lean @@ -0,0 +1,155 @@ +/- +Copyright (c) 2024 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ + +import Mathlib.CategoryTheory.Triangulated.Opposite +import Mathlib.CategoryTheory.Shift.ShiftedHom + +/-! Shifted morphisms in the opposite category + +If `C` is a category equipped with a shift by `ℤ`, `X` and `Y` are objects +of `C`, and `n : ℤ`, we define a bijection +`ShiftedHom.opEquiv : ShiftedHom X Y n ≃ ShiftedHom (Opposite.op Y) (Opposite.op X) n`. +We also introduce `ShiftedHom.opEquiv'` which produces a bijection +`ShiftedHom X Y a' ≃ (Opposite.op (Y⟦a⟧) ⟶ (Opposite.op X)⟦n⟧)` when `n + a = a'`. +The compatibilities that are obtained shall be used in order to study +the homological functor `preadditiveYoneda.obj B : Cᵒᵖ ⥤ Type _` when `B` is an object +in a pretriangulated category `C`. + +-/ + +namespace CategoryTheory + +open Category Pretriangulated.Opposite Pretriangulated + +variable {C : Type*} [Category C] [HasShift C ℤ] {X Y Z : C} + +namespace ShiftedHom + +/-- The bijection `ShiftedHom X Y n ≃ ShiftedHom (Opposite.op Y) (Opposite.op X) n` when +`n : ℤ`, and `X` and `Y` are objects of a category equipped with a shift by `ℤ`. -/ +noncomputable def opEquiv (n : ℤ) : + ShiftedHom X Y n ≃ ShiftedHom (Opposite.op Y) (Opposite.op X) n := + Quiver.Hom.opEquiv.trans + ((opShiftFunctorEquivalence C n).symm.toAdjunction.homEquiv (Opposite.op Y) (Opposite.op X)) + +lemma opEquiv_symm_apply {n : ℤ} (f : ShiftedHom (Opposite.op Y) (Opposite.op X) n) : + (opEquiv n).symm f = + ((opShiftFunctorEquivalence C n).unitIso.inv.app (Opposite.op X)).unop ≫ f.unop⟦n⟧' := by + rfl + +lemma opEquiv_symm_apply_comp {X Y : C} {a : ℤ} + (f : ShiftedHom (Opposite.op X) (Opposite.op Y) a) {b : ℤ} {Z : C} + (z : ShiftedHom X Z b) {c : ℤ} (h : b + a = c) : + ((ShiftedHom.opEquiv a).symm f).comp z h = + (ShiftedHom.opEquiv a).symm (z.op ≫ f) ≫ + (shiftFunctorAdd' C b a c h).inv.app Z := by + rw [ShiftedHom.opEquiv_symm_apply, ShiftedHom.opEquiv_symm_apply, + ShiftedHom.comp] + dsimp + simp only [assoc, Functor.map_comp] + +lemma opEquiv_symm_comp {a b : ℤ} + (f : ShiftedHom (Opposite.op Z) (Opposite.op Y) a) + (g : ShiftedHom (Opposite.op Y) (Opposite.op X) b) + {c : ℤ} (h : b + a = c) : + (opEquiv _).symm (f.comp g h) = + ((opEquiv _).symm g).comp ((opEquiv _).symm f) (by omega) := by + rw [opEquiv_symm_apply, opEquiv_symm_apply, + opShiftFunctorEquivalence_unitIso_inv_app_eq _ _ _ _ (show a + b = c by omega), comp, comp] + dsimp + rw [assoc, assoc, assoc, assoc, ← Functor.map_comp, ← unop_comp_assoc, + Iso.inv_hom_id_app] + dsimp + rw [assoc, id_comp, Functor.map_comp, ← NatTrans.naturality_assoc, + ← NatTrans.naturality, opEquiv_symm_apply] + dsimp + rw [← Functor.map_comp_assoc, ← Functor.map_comp_assoc, + ← Functor.map_comp_assoc] + rw [← unop_comp_assoc] + erw [← NatTrans.naturality] + rfl + +/-- The bijection `ShiftedHom X Y a' ≃ (Opposite.op (Y⟦a⟧) ⟶ (Opposite.op X)⟦n⟧)` +when integers `n`, `a` and `a'` satisfy `n + a = a'`, and `X` and `Y` are objects +of a category equipped with a shift by `ℤ`. -/ +noncomputable def opEquiv' (n a a' : ℤ) (h : n + a = a') : + ShiftedHom X Y a' ≃ (Opposite.op (Y⟦a⟧) ⟶ (Opposite.op X)⟦n⟧) := + ((shiftFunctorAdd' C a n a' (by omega)).symm.app Y).homToEquiv.symm.trans (opEquiv n) + +lemma opEquiv'_symm_apply {n a : ℤ} (f : Opposite.op (Y⟦a⟧) ⟶ (Opposite.op X)⟦n⟧) + (a' : ℤ) (h : n + a = a') : + (opEquiv' n a a' h).symm f = + (opEquiv n).symm f ≫ (shiftFunctorAdd' C a n a' (by omega)).inv.app _ := + rfl + +lemma opEquiv'_apply {a' : ℤ} (f : ShiftedHom X Y a') (n a : ℤ) (h : n + a = a') : + opEquiv' n a a' h f = + opEquiv n (f ≫ (shiftFunctorAdd' C a n a' (by omega)).hom.app Y) := by + rfl + +lemma opEquiv'_symm_op_opShiftFunctorEquivalence_counitIso_inv_app_op_shift + {n m : ℤ} (f : ShiftedHom X Y n) (g : ShiftedHom Y Z m) + (q : ℤ) (hq : n + m = q) : + (opEquiv' n m q hq).symm + (g.op ≫ (opShiftFunctorEquivalence C n).counitIso.inv.app _ ≫ f.op⟦n⟧') = + f.comp g (by omega) := by + rw [opEquiv'_symm_apply, opEquiv_symm_apply] + dsimp [comp] + apply Quiver.Hom.op_inj + simp only [assoc, Functor.map_comp, op_comp, Quiver.Hom.op_unop, + opShiftFunctorEquivalence_unitIso_inv_naturality] + erw [(opShiftFunctorEquivalence C n).inverse_counitInv_comp_assoc (Opposite.op Y)] + +lemma opEquiv'_symm_comp (f : Y ⟶ X) {n a : ℤ} (x : Opposite.op (Z⟦a⟧) ⟶ (Opposite.op X⟦n⟧)) + (a' : ℤ) (h : n + a = a') : + (opEquiv' n a a' h).symm (x ≫ f.op⟦n⟧') = f ≫ (opEquiv' n a a' h).symm x := + Quiver.Hom.op_inj (by simp [opEquiv'_symm_apply, opEquiv_symm_apply]) + +lemma opEquiv'_zero_add_symm (a : ℤ) (f : Opposite.op (Y⟦a⟧) ⟶ (Opposite.op X)⟦(0 : ℤ)⟧) : + (opEquiv' 0 a a (zero_add a)).symm f = + ((shiftFunctorZero Cᵒᵖ ℤ).hom.app _).unop ≫ f.unop := by + simp [opEquiv'_symm_apply, opEquiv_symm_apply, shiftFunctorAdd'_add_zero, + opShiftFunctorEquivalence_zero_unitIso_inv_app] + +lemma opEquiv'_add_symm (n m a a' a'' : ℤ) (ha' : n + a = a') (ha'' : m + a' = a'') + (x : (Opposite.op (Y⟦a⟧) ⟶ (Opposite.op X)⟦m + n⟧)) : + (opEquiv' (m + n) a a'' (by omega)).symm x = + (opEquiv' m a' a'' ha'').symm ((opEquiv' n a a' ha').symm + (x ≫ (shiftFunctorAdd Cᵒᵖ m n).hom.app _)).op := by + simp only [opEquiv'_symm_apply, opEquiv_symm_apply, + opShiftFunctorEquivalence_unitIso_inv_app_eq _ _ _ _ (add_comm n m)] + dsimp + simp only [assoc, Functor.map_comp, ← shiftFunctorAdd'_eq_shiftFunctorAdd, + ← NatTrans.naturality_assoc, + shiftFunctorAdd'_assoc_inv_app a n m a' (m + n) a'' (by omega) (by omega) (by omega)] + rfl + +section Preadditive + +variable [Preadditive C] [∀ (n : ℤ), (shiftFunctor C n).Additive] + +@[simp] +lemma opEquiv_symm_add {n : ℤ} (x y : ShiftedHom (Opposite.op Y) (Opposite.op X) n) : + (opEquiv n).symm (x + y) = (opEquiv n).symm x + (opEquiv n).symm y := by + dsimp [opEquiv_symm_apply] + rw [← Preadditive.comp_add, ← Functor.map_add] + rfl + +@[simp] +lemma opEquiv'_symm_add {n a : ℤ} (x y : (Opposite.op (Y⟦a⟧) ⟶ (Opposite.op X)⟦n⟧)) + (a' : ℤ) (h : n + a = a') : + (opEquiv' n a a' h).symm (x + y) = + (opEquiv' n a a' h).symm x + (opEquiv' n a a' h).symm y := by + dsimp [opEquiv'] + erw [opEquiv_symm_add, Iso.homToEquiv_apply, Iso.homToEquiv_apply, Iso.homToEquiv_apply] + rw [Preadditive.add_comp] + rfl + +end Preadditive + +end ShiftedHom + +end CategoryTheory diff --git a/Mathlib/CategoryTheory/Triangulated/Opposite.lean b/Mathlib/CategoryTheory/Triangulated/Opposite.lean index 3b6c664323762..9e34b1cc2e37f 100644 --- a/Mathlib/CategoryTheory/Triangulated/Opposite.lean +++ b/Mathlib/CategoryTheory/Triangulated/Opposite.lean @@ -130,8 +130,7 @@ lemma shiftFunctor_op_map (n m : ℤ) (hnm : n + m = 0) {K L : Cᵒᵖ} (φ : K (shiftFunctorOpIso C n m hnm).inv.app L := (NatIso.naturality_2 (shiftFunctorOpIso C n m hnm) φ).symm -variable (C) - +variable (C) in /-- The autoequivalence `Cᵒᵖ ≌ Cᵒᵖ` whose functor is `shiftFunctor Cᵒᵖ n` and whose inverse functor is `(shiftFunctor C n).op`. Do not unfold the definitions of the unit and counit isomorphisms: the compatibilities they satisfy are stated as separate lemmas. -/ @@ -177,7 +176,57 @@ lemma opShiftFunctorEquivalence_counitIso_inv_naturality (n : ℤ) {X Y : Cᵒ (opShiftFunctorEquivalence C n).counitIso.inv.app X ≫ f.unop⟦n⟧'.op⟦n⟧' := (opShiftFunctorEquivalence C n).counitIso.inv.naturality f -variable {C} +lemma opShiftFunctorEquivalence_zero_unitIso_hom_app (X : Cᵒᵖ) : + (opShiftFunctorEquivalence C 0).unitIso.hom.app X = + ((shiftFunctorZero C ℤ).hom.app X.unop).op ≫ + (((shiftFunctorZero Cᵒᵖ ℤ).inv.app X).unop⟦(0 : ℤ)⟧').op := by + apply Quiver.Hom.unop_inj + dsimp [opShiftFunctorEquivalence] + rw [shiftFunctorZero_op_inv_app, unop_comp, Quiver.Hom.unop_op, Functor.map_comp, + shiftFunctorCompIsoId_zero_zero_hom_app, assoc] + +lemma opShiftFunctorEquivalence_zero_unitIso_inv_app (X : Cᵒᵖ) : + (opShiftFunctorEquivalence C 0).unitIso.inv.app X = + (((shiftFunctorZero Cᵒᵖ ℤ).hom.app X).unop⟦(0 : ℤ)⟧').op ≫ + ((shiftFunctorZero C ℤ).inv.app X.unop).op := by + apply Quiver.Hom.unop_inj + dsimp [opShiftFunctorEquivalence] + rw [shiftFunctorZero_op_hom_app, unop_comp, Quiver.Hom.unop_op, Functor.map_comp, + shiftFunctorCompIsoId_zero_zero_inv_app, assoc] + +lemma opShiftFunctorEquivalence_unitIso_hom_app_eq (X : Cᵒᵖ) (m n p : ℤ) (h : m + n = p) : + (opShiftFunctorEquivalence C p).unitIso.hom.app X = + (opShiftFunctorEquivalence C n).unitIso.hom.app X ≫ + (((opShiftFunctorEquivalence C m).unitIso.hom.app (X⟦n⟧)).unop⟦n⟧').op ≫ + ((shiftFunctorAdd' C m n p h).hom.app _).op ≫ + (((shiftFunctorAdd' Cᵒᵖ n m p (by omega)).inv.app X).unop⟦p⟧').op := by + dsimp [opShiftFunctorEquivalence] + simp only [shiftFunctorAdd'_op_inv_app _ n m p (by omega) _ _ _ (add_neg_cancel n) + (add_neg_cancel m) (add_neg_cancel p), shiftFunctor_op_map _ _ (add_neg_cancel m), + Category.assoc, Iso.inv_hom_id_app_assoc] + erw [Functor.map_id, Functor.map_id, Functor.map_id, Functor.map_id, + id_comp, id_comp, id_comp, comp_id, comp_id] + dsimp + rw [comp_id, shiftFunctorCompIsoId_add'_hom_app _ _ _ _ _ _ + (neg_add_cancel m) (neg_add_cancel n) (neg_add_cancel p) h] + dsimp + rw [Category.assoc, Category.assoc] + rfl + +lemma opShiftFunctorEquivalence_unitIso_inv_app_eq (X : Cᵒᵖ) (m n p : ℤ) (h : m + n = p) : + (opShiftFunctorEquivalence C p).unitIso.inv.app X = + (((shiftFunctorAdd' Cᵒᵖ n m p (by omega)).hom.app X).unop⟦p⟧').op ≫ + ((shiftFunctorAdd' C m n p h).inv.app _).op ≫ + (((opShiftFunctorEquivalence C m).unitIso.inv.app (X⟦n⟧)).unop⟦n⟧').op ≫ + (opShiftFunctorEquivalence C n).unitIso.inv.app X := by + rw [← cancel_mono ((opShiftFunctorEquivalence C p).unitIso.hom.app X), Iso.inv_hom_id_app, + opShiftFunctorEquivalence_unitIso_hom_app_eq _ _ _ _ h, + Category.assoc, Category.assoc, Category.assoc, Iso.inv_hom_id_app_assoc] + apply Quiver.Hom.unop_inj + dsimp + simp only [Category.assoc, ← Functor.map_comp_assoc, Iso.hom_inv_id_app_assoc, + ← unop_comp, Iso.inv_hom_id_app, Functor.comp_obj, Functor.op_obj, unop_id, + Functor.map_id, id_comp, ← Functor.map_comp, Iso.hom_inv_id_app] lemma shift_unop_opShiftFunctorEquivalence_counitIso_inv_app (X : Cᵒᵖ) (n : ℤ) : ((opShiftFunctorEquivalence C n).counitIso.inv.app X).unop⟦n⟧' = diff --git a/Mathlib/CategoryTheory/Triangulated/Yoneda.lean b/Mathlib/CategoryTheory/Triangulated/Yoneda.lean index 39ec3c60ebdb1..e136effd690d8 100644 --- a/Mathlib/CategoryTheory/Triangulated/Yoneda.lean +++ b/Mathlib/CategoryTheory/Triangulated/Yoneda.lean @@ -5,6 +5,7 @@ Authors: Joël Riou -/ import Mathlib.Algebra.Homology.ShortComplex.Ab import Mathlib.CategoryTheory.Preadditive.Yoneda.Basic +import Mathlib.CategoryTheory.Shift.ShiftedHomOpposite import Mathlib.CategoryTheory.Triangulated.HomologicalFunctor import Mathlib.CategoryTheory.Triangulated.Opposite @@ -23,7 +24,7 @@ variable {C : Type*} [Category C] [Preadditive C] [HasShift C ℤ] namespace CategoryTheory -open Limits Pretriangulated.Opposite +open Limits Opposite Pretriangulated.Opposite namespace Pretriangulated @@ -62,6 +63,43 @@ lemma preadditiveCoyoneda_homologySequenceδ_apply x ≫ T.mor₃⟦n₀⟧' ≫ (shiftFunctorAdd' C 1 n₀ n₁ (by omega)).inv.app _ := by apply Category.assoc +section + +variable [∀ (n : ℤ), (shiftFunctor C n).Additive] + +noncomputable instance (B : C) : (preadditiveYoneda.obj B).ShiftSequence ℤ where + sequence n := preadditiveYoneda.obj (B⟦n⟧) + isoZero := preadditiveYoneda.mapIso ((shiftFunctorZero C ℤ).app B) + shiftIso n a a' h := NatIso.ofComponents (fun A ↦ AddEquiv.toAddCommGrpIso + { toEquiv := Quiver.Hom.opEquiv.trans (ShiftedHom.opEquiv' n a a' h).symm + map_add' := fun _ _ ↦ ShiftedHom.opEquiv'_symm_add _ _ _ h }) + (by intros; ext; apply ShiftedHom.opEquiv'_symm_comp _ _ _ h) + shiftIso_zero a := by ext; apply ShiftedHom.opEquiv'_zero_add_symm + shiftIso_add n m a a' a'' ha' ha'' := by + ext _ x + exact ShiftedHom.opEquiv'_add_symm n m a a' a'' ha' ha'' x.op + +lemma preadditiveYoneda_shiftMap_apply (B : C) {X Y : Cᵒᵖ} (n : ℤ) (f : X ⟶ Y⟦n⟧) + (a a' : ℤ) (h : n + a = a') (z : X.unop ⟶ B⟦a⟧) : + (preadditiveYoneda.obj B).shiftMap f a a' h z = + ((ShiftedHom.opEquiv _).symm f).comp z (show a + n = a' by omega) := by + symm + apply ShiftedHom.opEquiv_symm_apply_comp + +lemma preadditiveYoneda_homologySequenceδ_apply + (T : Triangle C) (n₀ n₁ : ℤ) (h : n₀ + 1 = n₁) {B : C} (x : T.obj₁ ⟶ B⟦n₀⟧) : + (preadditiveYoneda.obj B).homologySequenceδ + ((triangleOpEquivalence _).functor.obj (op T)) n₀ n₁ h x = + T.mor₃ ≫ x⟦(1 : ℤ)⟧' ≫ (shiftFunctorAdd' C n₀ 1 n₁ h).inv.app B := by + simp only [Functor.homologySequenceδ, preadditiveYoneda_shiftMap_apply, + ShiftedHom.comp, ← Category.assoc] + congr 2 + apply (ShiftedHom.opEquiv _).injective + rw [Equiv.apply_symm_apply] + rfl + +end + end Pretriangulated end CategoryTheory diff --git a/Mathlib/Combinatorics/Quiver/Basic.lean b/Mathlib/Combinatorics/Quiver/Basic.lean index f0f77f70aa4b7..78d2fbdeef06a 100644 --- a/Mathlib/Combinatorics/Quiver/Basic.lean +++ b/Mathlib/Combinatorics/Quiver/Basic.lean @@ -137,6 +137,15 @@ def Hom.op {V} [Quiver V] {X Y : V} (f : X ⟶ Y) : op Y ⟶ op X := ⟨f⟩ /-- Given an arrow in `Vᵒᵖ`, we can take the "unopposite" back in `V`. -/ def Hom.unop {V} [Quiver V] {X Y : Vᵒᵖ} (f : X ⟶ Y) : unop Y ⟶ unop X := Opposite.unop f +/-- The bijection `(X ⟶ Y) ≃ (op Y ⟶ op X)`. -/ +@[simps] +def Hom.opEquiv {V} [Quiver V] {X Y : V} : + (X ⟶ Y) ≃ (Opposite.op Y ⟶ Opposite.op X) where + toFun := Opposite.op + invFun := Opposite.unop + left_inv _ := rfl + right_inv _ := rfl + /-- A type synonym for a quiver with no arrows. -/ -- Porting note(#5171): this linter isn't ported yet. -- @[nolint has_nonempty_instance] From cbcf4ea1cb4620e09549662684cb12de35d81299 Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Thu, 17 Oct 2024 15:41:50 +0000 Subject: [PATCH 100/505] feat(Analysis/Normed/Group/Ultra): triangle ineq for tsum & tprod (#17584) --- Mathlib/Analysis/Normed/Group/Ultra.lean | 42 ++++++++++++++++++++++++ 1 file changed, 42 insertions(+) diff --git a/Mathlib/Analysis/Normed/Group/Ultra.lean b/Mathlib/Analysis/Normed/Group/Ultra.lean index 8e83b3535e430..2b8c5f341691f 100644 --- a/Mathlib/Analysis/Normed/Group/Ultra.lean +++ b/Mathlib/Analysis/Normed/Group/Ultra.lean @@ -6,6 +6,8 @@ Authors: Yakov Pechersky, David Loeffler import Mathlib.Analysis.Normed.Group.Uniform import Mathlib.Topology.Algebra.Nonarchimedean.Basic import Mathlib.Topology.MetricSpace.Ultra.Basic +import Mathlib.Topology.Algebra.InfiniteSum.Group +import Mathlib.Topology.Algebra.Order.LiminfLimsup /-! # Ultrametric norms @@ -31,6 +33,8 @@ ultrametric, nonarchimedean -/ open Metric NNReal +open scoped BigOperators + namespace IsUltrametricDist section Group @@ -235,6 +239,44 @@ lemma norm_prod_le_of_forall_le_of_nonneg {s : Finset ι} {f : ι → M} {C : lift C to NNReal using h_nonneg exact nnnorm_prod_le_of_forall_le hC +@[to_additive] +lemma norm_tprod_le (f : ι → M) : ‖∏' i, f i‖ ≤ ⨆ i, ‖f i‖ := by + rcases isEmpty_or_nonempty ι with hι | hι + · -- Silly case #1 : the index type is empty + simp only [tprod_empty, norm_one', Real.iSup_of_isEmpty, le_refl] + by_cases h : Multipliable f; swap + · -- Silly case #2 : the product is divergent + rw [tprod_eq_one_of_not_multipliable h, norm_one'] + by_cases h_bd : BddAbove (Set.range fun i ↦ ‖f i‖) + · exact le_ciSup_of_le h_bd hι.some (norm_nonneg' _) + · rw [Real.iSup_of_not_bddAbove h_bd] + -- now the interesting case + have h_bd : BddAbove (Set.range fun i ↦ ‖f i‖) := + h.tendsto_cofinite_one.norm'.bddAbove_range_of_cofinite + refine le_of_tendsto' h.hasProd.norm' (fun s ↦ norm_prod_le_of_forall_le_of_nonneg ?_ ?_) + · exact le_ciSup_of_le h_bd hι.some (norm_nonneg' _) + · exact fun i _ ↦ le_ciSup h_bd i + +@[to_additive] +lemma nnnorm_tprod_le (f : ι → M) : ‖∏' i, f i‖₊ ≤ ⨆ i, ‖f i‖₊ := by + simpa only [← NNReal.coe_le_coe, coe_nnnorm', coe_iSup] using norm_tprod_le f + +@[to_additive] +lemma norm_tprod_le_of_forall_le [Nonempty ι] {f : ι → M} {C : ℝ} (h : ∀ i, ‖f i‖ ≤ C) : + ‖∏' i, f i‖ ≤ C := + (norm_tprod_le f).trans (ciSup_le h) + +@[to_additive] +lemma norm_tprod_le_of_forall_le_of_nonneg {f : ι → M} {C : ℝ} (hC : 0 ≤ C) (h : ∀ i, ‖f i‖ ≤ C) : + ‖∏' i, f i‖ ≤ C := by + rcases isEmpty_or_nonempty ι + · simpa only [tprod_empty, norm_one'] using hC + · exact norm_tprod_le_of_forall_le h + +@[to_additive] +lemma nnnorm_tprod_le_of_forall_le {f : ι → M} {C : ℝ≥0} (h : ∀ i, ‖f i‖₊ ≤ C) : ‖∏' i, f i‖₊ ≤ C := + (nnnorm_tprod_le f).trans (ciSup_le' h) + end CommGroup end IsUltrametricDist From 0456b9a130212be2501255a28f73dbc146e05de2 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 17 Oct 2024 15:41:52 +0000 Subject: [PATCH 101/505] feat: `#s` as notation for `Finset.card s` (#17646) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit I believe it is time that we finally introduce this notation. This makes serious finset calculations significantly nicer (see eg #5297) and is standard notation. `|s|` would be an alternative, but absolute value elaboration is very complicated like that, so I would rather not, and `#{x | p x}` looks nicer than `|{x | p x}|` (although in return `|s ∩ t|` looks nicer than `#(s ∩ t)`). I am scoping the notation to the `Finset` namespace to avoid clashes with `Cardinal.mk`, which also uses that notation. --- Mathlib/Data/Finset/Card.lean | 316 +++++++++++++++++----------------- 1 file changed, 161 insertions(+), 155 deletions(-) diff --git a/Mathlib/Data/Finset/Card.lean b/Mathlib/Data/Finset/Card.lean index 653bbfc7a2fe2..376e7fe70747f 100644 --- a/Mathlib/Data/Finset/Card.lean +++ b/Mathlib/Data/Finset/Card.lean @@ -12,7 +12,7 @@ This defines the cardinality of a `Finset` and provides induction principles for ## Main declarations -* `Finset.card`: `s.card : ℕ` returns the cardinality of `s : Finset α`. +* `Finset.card`: `#s : ℕ` returns the cardinality of `s : Finset α`. ### Induction principles @@ -36,52 +36,56 @@ namespace Finset variable {s t : Finset α} {a b : α} -/-- `s.card` is the number of elements of `s`, aka its cardinality. -/ +/-- `s.card` is the number of elements of `s`, aka its cardinality. + +The notation `#s` can be accessed in the `Finset` locale. -/ def card (s : Finset α) : ℕ := Multiset.card s.1 -theorem card_def (s : Finset α) : s.card = Multiset.card s.1 := +@[inherit_doc] scoped prefix:arg "#" => Finset.card + +theorem card_def (s : Finset α) : #s = Multiset.card s.1 := rfl -@[simp] lemma card_val (s : Finset α) : Multiset.card s.1 = s.card := rfl +@[simp] lemma card_val (s : Finset α) : Multiset.card s.1 = #s := rfl @[simp] -theorem card_mk {m nodup} : (⟨m, nodup⟩ : Finset α).card = Multiset.card m := +theorem card_mk {m nodup} : #(⟨m, nodup⟩ : Finset α) = Multiset.card m := rfl @[simp] -theorem card_empty : card (∅ : Finset α) = 0 := +theorem card_empty : #(∅ : Finset α) = 0 := rfl @[gcongr] -theorem card_le_card : s ⊆ t → s.card ≤ t.card := +theorem card_le_card : s ⊆ t → #s ≤ #t := Multiset.card_le_card ∘ val_le_iff.mpr @[mono] theorem card_mono : Monotone (@card α) := by apply card_le_card -@[simp] lemma card_eq_zero : s.card = 0 ↔ s = ∅ := Multiset.card_eq_zero.trans val_eq_zero -lemma card_ne_zero : s.card ≠ 0 ↔ s.Nonempty := card_eq_zero.ne.trans nonempty_iff_ne_empty.symm -@[simp] lemma card_pos : 0 < s.card ↔ s.Nonempty := Nat.pos_iff_ne_zero.trans card_ne_zero -@[simp] lemma one_le_card : 1 ≤ s.card ↔ s.Nonempty := card_pos +@[simp] lemma card_eq_zero : #s = 0 ↔ s = ∅ := Multiset.card_eq_zero.trans val_eq_zero +lemma card_ne_zero : #s ≠ 0 ↔ s.Nonempty := card_eq_zero.ne.trans nonempty_iff_ne_empty.symm +@[simp] lemma card_pos : 0 < #s ↔ s.Nonempty := Nat.pos_iff_ne_zero.trans card_ne_zero +@[simp] lemma one_le_card : 1 ≤ #s ↔ s.Nonempty := card_pos alias ⟨_, Nonempty.card_pos⟩ := card_pos alias ⟨_, Nonempty.card_ne_zero⟩ := card_ne_zero -theorem card_ne_zero_of_mem (h : a ∈ s) : s.card ≠ 0 := +theorem card_ne_zero_of_mem (h : a ∈ s) : #s ≠ 0 := (not_congr card_eq_zero).2 <| ne_empty_of_mem h @[simp] -theorem card_singleton (a : α) : card ({a} : Finset α) = 1 := +theorem card_singleton (a : α) : #{a} = 1 := Multiset.card_singleton _ -theorem card_singleton_inter [DecidableEq α] : ({a} ∩ s).card ≤ 1 := by +theorem card_singleton_inter [DecidableEq α] : #({a} ∩ s) ≤ 1 := by cases' Finset.decidableMem a s with h h · simp [Finset.singleton_inter_of_not_mem h] · simp [Finset.singleton_inter_of_mem h] @[simp] -theorem card_cons (h : a ∉ s) : (s.cons a h).card = s.card + 1 := +theorem card_cons (h : a ∉ s) : #(s.cons a h) = #s + 1 := Multiset.card_cons _ _ section InsertErase @@ -89,12 +93,12 @@ section InsertErase variable [DecidableEq α] @[simp] -theorem card_insert_of_not_mem (h : a ∉ s) : (insert a s).card = s.card + 1 := by +theorem card_insert_of_not_mem (h : a ∉ s) : #(insert a s) = #s + 1 := by rw [← cons_eq_insert _ _ h, card_cons] -theorem card_insert_of_mem (h : a ∈ s) : card (insert a s) = s.card := by rw [insert_eq_of_mem h] +theorem card_insert_of_mem (h : a ∈ s) : #(insert a s) = #s := by rw [insert_eq_of_mem h] -theorem card_insert_le (a : α) (s : Finset α) : card (insert a s) ≤ s.card + 1 := by +theorem card_insert_le (a : α) (s : Finset α) : #(insert a s) ≤ #s + 1 := by by_cases h : a ∈ s · rw [insert_eq_of_mem h] exact Nat.le_succ _ @@ -104,43 +108,43 @@ section variable {a b c d e f : α} -theorem card_le_two : card {a, b} ≤ 2 := card_insert_le _ _ +theorem card_le_two : #{a, b} ≤ 2 := card_insert_le _ _ -theorem card_le_three : card {a, b, c} ≤ 3 := +theorem card_le_three : #{a, b, c} ≤ 3 := (card_insert_le _ _).trans (Nat.succ_le_succ card_le_two) -theorem card_le_four : card {a, b, c, d} ≤ 4 := +theorem card_le_four : #{a, b, c, d} ≤ 4 := (card_insert_le _ _).trans (Nat.succ_le_succ card_le_three) -theorem card_le_five : card {a, b, c, d, e} ≤ 5 := +theorem card_le_five : #{a, b, c, d, e} ≤ 5 := (card_insert_le _ _).trans (Nat.succ_le_succ card_le_four) -theorem card_le_six : card {a, b, c, d, e, f} ≤ 6 := +theorem card_le_six : #{a, b, c, d, e, f} ≤ 6 := (card_insert_le _ _).trans (Nat.succ_le_succ card_le_five) end /-- If `a ∈ s` is known, see also `Finset.card_insert_of_mem` and `Finset.card_insert_of_not_mem`. -/ -theorem card_insert_eq_ite : card (insert a s) = if a ∈ s then s.card else s.card + 1 := by +theorem card_insert_eq_ite : #(insert a s) = if a ∈ s then #s else #s + 1 := by by_cases h : a ∈ s · rw [card_insert_of_mem h, if_pos h] · rw [card_insert_of_not_mem h, if_neg h] @[simp] -theorem card_pair_eq_one_or_two : ({a,b} : Finset α).card = 1 ∨ ({a,b} : Finset α).card = 2 := by +theorem card_pair_eq_one_or_two : #{a, b} = 1 ∨ #{a, b} = 2 := by simp [card_insert_eq_ite] tauto @[simp] -theorem card_pair (h : a ≠ b) : ({a, b} : Finset α).card = 2 := by +theorem card_pair (h : a ≠ b) : #{a, b} = 2 := by rw [card_insert_of_not_mem (not_mem_singleton.2 h), card_singleton] @[deprecated (since := "2024-01-04")] alias card_doubleton := Finset.card_pair /-- $\#(s \setminus \{a\}) = \#s - 1$ if $a \in s$. -/ @[simp] -theorem card_erase_of_mem : a ∈ s → (s.erase a).card = s.card - 1 := +theorem card_erase_of_mem : a ∈ s → #(s.erase a) = #s - 1 := Multiset.card_erase_of_mem /-- $\#(s \setminus \{a\}) = \#s - 1$ if $a \in s$. @@ -148,55 +152,57 @@ theorem card_erase_of_mem : a ∈ s → (s.erase a).card = s.card - 1 := so that we don't have to work with `ℕ`-subtraction. -/ @[simp] theorem cast_card_erase_of_mem {R} [AddGroupWithOne R] {s : Finset α} (hs : a ∈ s) : - ((s.erase a).card : R) = s.card - 1 := by + (#(s.erase a) : R) = #s - 1 := by rw [card_erase_of_mem hs, Nat.cast_sub, Nat.cast_one] rw [Nat.add_one_le_iff, Finset.card_pos] exact ⟨a, hs⟩ @[simp] -theorem card_erase_add_one : a ∈ s → (s.erase a).card + 1 = s.card := +theorem card_erase_add_one : a ∈ s → #(s.erase a) + 1 = #s := Multiset.card_erase_add_one -theorem card_erase_lt_of_mem : a ∈ s → (s.erase a).card < s.card := +theorem card_erase_lt_of_mem : a ∈ s → #(s.erase a) < #s := Multiset.card_erase_lt_of_mem -theorem card_erase_le : (s.erase a).card ≤ s.card := +theorem card_erase_le : #(s.erase a) ≤ #s := Multiset.card_erase_le -theorem pred_card_le_card_erase : s.card - 1 ≤ (s.erase a).card := by +theorem pred_card_le_card_erase : #s - 1 ≤ #(s.erase a) := by by_cases h : a ∈ s · exact (card_erase_of_mem h).ge · rw [erase_eq_of_not_mem h] exact Nat.sub_le _ _ /-- If `a ∈ s` is known, see also `Finset.card_erase_of_mem` and `Finset.erase_eq_of_not_mem`. -/ -theorem card_erase_eq_ite : (s.erase a).card = if a ∈ s then s.card - 1 else s.card := +theorem card_erase_eq_ite : #(s.erase a) = if a ∈ s then #s - 1 else #s := Multiset.card_erase_eq_ite end InsertErase @[simp] -theorem card_range (n : ℕ) : (range n).card = n := +theorem card_range (n : ℕ) : #(range n) = n := Multiset.card_range n @[simp] -theorem card_attach : s.attach.card = s.card := +theorem card_attach : #s.attach = #s := Multiset.card_attach end Finset +open scoped Finset + section ToMLListultiset variable [DecidableEq α] (m : Multiset α) (l : List α) -theorem Multiset.card_toFinset : m.toFinset.card = Multiset.card m.dedup := +theorem Multiset.card_toFinset : #m.toFinset = Multiset.card m.dedup := rfl -theorem Multiset.toFinset_card_le : m.toFinset.card ≤ Multiset.card m := +theorem Multiset.toFinset_card_le : #m.toFinset ≤ Multiset.card m := card_le_card <| dedup_le _ theorem Multiset.toFinset_card_of_nodup {m : Multiset α} (h : m.Nodup) : - m.toFinset.card = Multiset.card m := + #m.toFinset = Multiset.card m := congr_arg card <| Multiset.dedup_eq_self.mpr h theorem Multiset.dedup_card_eq_card_iff_nodup {m : Multiset α} : @@ -204,15 +210,15 @@ theorem Multiset.dedup_card_eq_card_iff_nodup {m : Multiset α} : .trans ⟨fun h ↦ eq_of_le_of_card_le (dedup_le m) h.ge, congr_arg _⟩ dedup_eq_self theorem Multiset.toFinset_card_eq_card_iff_nodup {m : Multiset α} : - m.toFinset.card = card m ↔ m.Nodup := dedup_card_eq_card_iff_nodup + #m.toFinset = card m ↔ m.Nodup := dedup_card_eq_card_iff_nodup -theorem List.card_toFinset : l.toFinset.card = l.dedup.length := +theorem List.card_toFinset : #l.toFinset = l.dedup.length := rfl -theorem List.toFinset_card_le : l.toFinset.card ≤ l.length := +theorem List.toFinset_card_le : #l.toFinset ≤ l.length := Multiset.toFinset_card_le ⟦l⟧ -theorem List.toFinset_card_of_nodup {l : List α} (h : l.Nodup) : l.toFinset.card = l.length := +theorem List.toFinset_card_of_nodup {l : List α} (h : l.Nodup) : #l.toFinset = l.length := Multiset.toFinset_card_of_nodup h end ToMLListultiset @@ -222,16 +228,16 @@ namespace Finset variable {s t u : Finset α} {f : α → β} {n : ℕ} @[simp] -theorem length_toList (s : Finset α) : s.toList.length = s.card := by +theorem length_toList (s : Finset α) : s.toList.length = #s := by rw [toList, ← Multiset.coe_card, Multiset.coe_toList, card_def] -theorem card_image_le [DecidableEq β] : (s.image f).card ≤ s.card := by +theorem card_image_le [DecidableEq β] : #(s.image f) ≤ #s := by simpa only [card_map] using (s.1.map f).toFinset_card_le -theorem card_image_of_injOn [DecidableEq β] (H : Set.InjOn f s) : (s.image f).card = s.card := by +theorem card_image_of_injOn [DecidableEq β] (H : Set.InjOn f s) : #(s.image f) = #s := by simp only [card, image_val_of_injOn H, card_map] -theorem injOn_of_card_image_eq [DecidableEq β] (H : (s.image f).card = s.card) : Set.InjOn f s := by +theorem injOn_of_card_image_eq [DecidableEq β] (H : #(s.image f) = #s) : Set.InjOn f s := by rw [card_def, card_def, image, toFinset] at H dsimp only at H have : (s.1.map f).dedup = s.1.map f := by @@ -240,58 +246,58 @@ theorem injOn_of_card_image_eq [DecidableEq β] (H : (s.image f).card = s.card) rw [Multiset.dedup_eq_self] at this exact inj_on_of_nodup_map this -theorem card_image_iff [DecidableEq β] : (s.image f).card = s.card ↔ Set.InjOn f s := +theorem card_image_iff [DecidableEq β] : #(s.image f) = #s ↔ Set.InjOn f s := ⟨injOn_of_card_image_eq, card_image_of_injOn⟩ theorem card_image_of_injective [DecidableEq β] (s : Finset α) (H : Injective f) : - (s.image f).card = s.card := + #(s.image f) = #s := card_image_of_injOn fun _ _ _ _ h => H h theorem fiber_card_ne_zero_iff_mem_image (s : Finset α) (f : α → β) [DecidableEq β] (y : β) : - (s.filter fun x => f x = y).card ≠ 0 ↔ y ∈ s.image f := by + #(s.filter fun x ↦ f x = y) ≠ 0 ↔ y ∈ s.image f := by rw [← Nat.pos_iff_ne_zero, card_pos, fiber_nonempty_iff_mem_image] lemma card_filter_le_iff (s : Finset α) (P : α → Prop) [DecidablePred P] (n : ℕ) : - (s.filter P).card ≤ n ↔ ∀ s' ⊆ s, n < s'.card → ∃ a ∈ s', ¬ P a := + #(s.filter P) ≤ n ↔ ∀ s' ⊆ s, n < #s' → ∃ a ∈ s', ¬ P a := (s.1.card_filter_le_iff P n).trans ⟨fun H s' hs' h ↦ H s'.1 (by aesop) h, fun H s' hs' h ↦ H ⟨s', nodup_of_le hs' s.2⟩ (fun _ hx ↦ subset_of_le hs' hx) h⟩ @[simp] -theorem card_map (f : α ↪ β) : (s.map f).card = s.card := +theorem card_map (f : α ↪ β) : #(s.map f) = #s := Multiset.card_map _ _ @[simp] theorem card_subtype (p : α → Prop) [DecidablePred p] (s : Finset α) : - (s.subtype p).card = (s.filter p).card := by simp [Finset.subtype] + #(s.subtype p) = #(s.filter p) := by simp [Finset.subtype] theorem card_filter_le (s : Finset α) (p : α → Prop) [DecidablePred p] : - (s.filter p).card ≤ s.card := + #(s.filter p) ≤ #s := card_le_card <| filter_subset _ _ -theorem eq_of_subset_of_card_le {s t : Finset α} (h : s ⊆ t) (h₂ : t.card ≤ s.card) : s = t := +theorem eq_of_subset_of_card_le {s t : Finset α} (h : s ⊆ t) (h₂ : #t ≤ #s) : s = t := eq_of_veq <| Multiset.eq_of_le_of_card_le (val_le_iff.mpr h) h₂ -theorem eq_of_superset_of_card_ge (hst : s ⊆ t) (hts : t.card ≤ s.card) : t = s := +theorem eq_of_superset_of_card_ge (hst : s ⊆ t) (hts : #t ≤ #s) : t = s := (eq_of_subset_of_card_le hst hts).symm -theorem subset_iff_eq_of_card_le (h : t.card ≤ s.card) : s ⊆ t ↔ s = t := +theorem subset_iff_eq_of_card_le (h : #t ≤ #s) : s ⊆ t ↔ s = t := ⟨fun hst => eq_of_subset_of_card_le hst h, Eq.subset'⟩ theorem map_eq_of_subset {f : α ↪ α} (hs : s.map f ⊆ s) : s.map f = s := eq_of_subset_of_card_le hs (card_map _).ge -theorem filter_card_eq {p : α → Prop} [DecidablePred p] (h : (s.filter p).card = s.card) (x : α) +theorem filter_card_eq {p : α → Prop} [DecidablePred p] (h : #(s.filter p) = #s) (x : α) (hx : x ∈ s) : p x := by rw [← eq_of_subset_of_card_le (s.filter_subset p) h.ge, mem_filter] at hx exact hx.2 -nonrec lemma card_lt_card (h : s ⊂ t) : s.card < t.card := card_lt_card <| val_lt_iff.2 h +nonrec lemma card_lt_card (h : s ⊂ t) : #s < #t := card_lt_card <| val_lt_iff.2 h lemma card_strictMono : StrictMono (card : Finset α → ℕ) := fun _ _ ↦ card_lt_card theorem card_eq_of_bijective (f : ∀ i, i < n → α) (hf : ∀ a ∈ s, ∃ i, ∃ h : i < n, f i h = a) (hf' : ∀ i (h : i < n), f i h ∈ s) - (f_inj : ∀ i j (hi : i < n) (hj : j < n), f i hi = f j hj → i = j) : s.card = n := by + (f_inj : ∀ i j (hi : i < n) (hj : j < n), f i hi = f j hj → i = j) : #s = n := by classical have : s = (range n).attach.image fun i => f i.1 (mem_range.1 i.2) := by ext a @@ -301,10 +307,10 @@ theorem card_eq_of_bijective (f : ∀ i, i < n → α) (hf : ∀ a ∈ s, ∃ i, · intro ha; obtain ⟨i, hi, rfl⟩ := hf a ha; use i, mem_range.2 hi · rintro ⟨i, hi, rfl⟩; apply hf' calc - s.card = ((range n).attach.image fun i => f i.1 (mem_range.1 i.2)).card := by rw [this] - _ = (range n).attach.card := ?_ - _ = (range n).card := card_attach - _ = n := card_range n + #s = #((range n).attach.image fun i => f i.1 (mem_range.1 i.2)) := by rw [this] + _ = #(range n).attach := ?_ + _ = #(range n) := card_attach + _ = n := card_range n apply card_image_of_injective intro ⟨i, hi⟩ ⟨j, hj⟩ eq exact Subtype.eq <| f_inj i j (mem_range.1 hi) (mem_range.1 hj) eq @@ -321,12 +327,12 @@ The difference with `Finset.card_nbij` is that the bijection is allowed to use m domain, rather than being a non-dependent function. -/ lemma card_bij (i : ∀ a ∈ s, β) (hi : ∀ a ha, i a ha ∈ t) (i_inj : ∀ a₁ ha₁ a₂ ha₂, i a₁ ha₁ = i a₂ ha₂ → a₁ = a₂) - (i_surj : ∀ b ∈ t, ∃ a ha, i a ha = b) : s.card = t.card := by + (i_surj : ∀ b ∈ t, ∃ a ha, i a ha = b) : #s = #t := by classical calc - s.card = s.attach.card := card_attach.symm - _ = (s.attach.image fun a : { a // a ∈ s } => i a.1 a.2).card := Eq.symm ?_ - _ = t.card := ?_ + #s = #s.attach := card_attach.symm + _ = #(s.attach.image fun a ↦ i a.1 a.2) := Eq.symm ?_ + _ = #t := ?_ · apply card_image_of_injective intro ⟨_, _⟩ ⟨_, _⟩ h simpa using i_inj _ _ _ _ h @@ -347,7 +353,7 @@ The difference with `Finset.card_nbij'` is that the bijection and its inverse ar membership of the domains, rather than being non-dependent functions. -/ lemma card_bij' (i : ∀ a ∈ s, β) (j : ∀ a ∈ t, α) (hi : ∀ a ha, i a ha ∈ t) (hj : ∀ a ha, j a ha ∈ s) (left_inv : ∀ a ha, j (i a ha) (hi a ha) = a) - (right_inv : ∀ a ha, i (j a ha) (hj a ha) = a) : s.card = t.card := by + (right_inv : ∀ a ha, i (j a ha) (hj a ha) = a) : #s = #t := by refine card_bij i hi (fun a1 h1 a2 h2 eq ↦ ?_) (fun b hb ↦ ⟨_, hj b hb, right_inv b hb⟩) rw [← left_inv a1 h1, ← left_inv a2 h2] simp only [eq] @@ -360,7 +366,7 @@ injection, rather than by an inverse function. The difference with `Finset.card_bij` is that the bijection is a non-dependent function, rather than being allowed to use membership of the domain. -/ lemma card_nbij (i : α → β) (hi : ∀ a ∈ s, i a ∈ t) (i_inj : (s : Set α).InjOn i) - (i_surj : (s : Set α).SurjOn i t) : s.card = t.card := + (i_surj : (s : Set α).SurjOn i t) : #s = #t := card_bij (fun a _ ↦ i a) hi i_inj (by simpa using i_surj) /-- Reorder a finset. @@ -374,35 +380,35 @@ functions, rather than being allowed to use membership of the domains. The difference with `Finset.card_equiv` is that bijectivity is only required to hold on the domains, rather than on the entire types. -/ lemma card_nbij' (i : α → β) (j : β → α) (hi : ∀ a ∈ s, i a ∈ t) (hj : ∀ a ∈ t, j a ∈ s) - (left_inv : ∀ a ∈ s, j (i a) = a) (right_inv : ∀ a ∈ t, i (j a) = a) : s.card = t.card := + (left_inv : ∀ a ∈ s, j (i a) = a) (right_inv : ∀ a ∈ t, i (j a) = a) : #s = #t := card_bij' (fun a _ ↦ i a) (fun b _ ↦ j b) hi hj left_inv right_inv /-- Specialization of `Finset.card_nbij'` that automatically fills in most arguments. See `Fintype.card_equiv` for the version where `s` and `t` are `univ`. -/ -lemma card_equiv (e : α ≃ β) (hst : ∀ i, i ∈ s ↔ e i ∈ t) : s.card = t.card := by +lemma card_equiv (e : α ≃ β) (hst : ∀ i, i ∈ s ↔ e i ∈ t) : #s = #t := by refine card_nbij' e e.symm ?_ ?_ ?_ ?_ <;> simp [hst] /-- Specialization of `Finset.card_nbij` that automatically fills in most arguments. See `Fintype.card_bijective` for the version where `s` and `t` are `univ`. -/ lemma card_bijective (e : α → β) (he : e.Bijective) (hst : ∀ i, i ∈ s ↔ e i ∈ t) : - s.card = t.card := card_equiv (.ofBijective e he) hst + #s = #t := card_equiv (.ofBijective e he) hst lemma card_le_card_of_injOn (f : α → β) (hf : ∀ a ∈ s, f a ∈ t) (f_inj : (s : Set α).InjOn f) : - s.card ≤ t.card := by + #s ≤ #t := by classical calc - s.card = (s.image f).card := (card_image_of_injOn f_inj).symm - _ ≤ t.card := card_le_card <| image_subset_iff.2 hf + #s = #(s.image f) := (card_image_of_injOn f_inj).symm + _ ≤ #t := card_le_card <| image_subset_iff.2 hf @[deprecated (since := "2024-06-01")] alias card_le_card_of_inj_on := card_le_card_of_injOn -lemma card_le_card_of_surjOn (f : α → β) (hf : Set.SurjOn f s t) : t.card ≤ s.card := by +lemma card_le_card_of_surjOn (f : α → β) (hf : Set.SurjOn f s t) : #t ≤ #s := by classical unfold Set.SurjOn at hf; exact (card_le_card (mod_cast hf)).trans card_image_le /-- If there are more pigeons than pigeonholes, then there are two pigeons in the same pigeonhole. -/ -theorem exists_ne_map_eq_of_card_lt_of_maps_to {t : Finset β} (hc : t.card < s.card) {f : α → β} +theorem exists_ne_map_eq_of_card_lt_of_maps_to {t : Finset β} (hc : #t < #s) {f : α → β} (hf : ∀ a ∈ s, f a ∈ t) : ∃ x ∈ s, ∃ y ∈ s, x ≠ y ∧ f x = f y := by classical by_contra! hz @@ -412,16 +418,16 @@ theorem exists_ne_map_eq_of_card_lt_of_maps_to {t : Finset β} (hc : t.card < s. exact hz x hx y hy lemma le_card_of_inj_on_range (f : ℕ → α) (hf : ∀ i < n, f i ∈ s) - (f_inj : ∀ i < n, ∀ j < n, f i = f j → i = j) : n ≤ s.card := + (f_inj : ∀ i < n, ∀ j < n, f i = f j → i = j) : n ≤ #s := calc - n = card (range n) := (card_range n).symm - _ ≤ s.card := card_le_card_of_injOn f (by simpa only [mem_range]) (by simpa) + n = #(range n) := (card_range n).symm + _ ≤ #s := card_le_card_of_injOn f (by simpa only [mem_range]) (by simpa) lemma surj_on_of_inj_on_of_card_le (f : ∀ a ∈ s, β) (hf : ∀ a ha, f a ha ∈ t) - (hinj : ∀ a₁ a₂ ha₁ ha₂, f a₁ ha₁ = f a₂ ha₂ → a₁ = a₂) (hst : t.card ≤ s.card) : + (hinj : ∀ a₁ a₂ ha₁ ha₂, f a₁ ha₁ = f a₂ ha₂ → a₁ = a₂) (hst : #t ≤ #s) : ∀ b ∈ t, ∃ a ha, b = f a ha := by classical - have h : (s.attach.image fun a : { a // a ∈ s } => f a a.prop).card = s.card := by + have h : #(s.attach.image fun a : s ↦ f a a.2) = #s := by rw [← @card_attach _ s, card_image_of_injective] intro ⟨_, _⟩ ⟨_, _⟩ h exact Subtype.eq <| hinj _ _ _ _ h @@ -431,7 +437,7 @@ lemma surj_on_of_inj_on_of_card_le (f : ∀ a ∈ s, β) (hf : ∀ a ha, f a ha exact fun b a ha hb ↦ ⟨a, ha, hb.symm⟩ theorem inj_on_of_surj_on_of_card_le (f : ∀ a ∈ s, β) (hf : ∀ a ha, f a ha ∈ t) - (hsurj : ∀ b ∈ t, ∃ a ha, f a ha = b) (hst : s.card ≤ t.card) ⦃a₁⦄ (ha₁ : a₁ ∈ s) ⦃a₂⦄ + (hsurj : ∀ b ∈ t, ∃ a ha, f a ha = b) (hst : #s ≤ #t) ⦃a₁⦄ (ha₁ : a₁ ∈ s) ⦃a₂⦄ (ha₂ : a₂ ∈ s) (ha₁a₂ : f a₁ ha₁ = f a₂ ha₂) : a₁ = a₂ := haveI : Inhabited { x // x ∈ s } := ⟨⟨a₁, ha₁⟩⟩ let f' : { x // x ∈ s } → { x // x ∈ t } := fun x => ⟨f x.1 x.2, hf x.1 x.2⟩ @@ -453,7 +459,7 @@ theorem inj_on_of_surj_on_of_card_le (f : ∀ a ∈ s, β) (hf : ∀ a ha, f a h end bij @[simp] -theorem card_disjUnion (s t : Finset α) (h) : (s.disjUnion t h).card = s.card + t.card := +theorem card_disjUnion (s t : Finset α) (h) : #(s.disjUnion t h) = #s + #t := Multiset.card_add _ _ /-! ### Lattice structure -/ @@ -464,23 +470,23 @@ section Lattice variable [DecidableEq α] theorem card_union_add_card_inter (s t : Finset α) : - (s ∪ t).card + (s ∩ t).card = s.card + t.card := + #(s ∪ t) + #(s ∩ t) = #s + #t := Finset.induction_on t (by simp) fun a r har h => by by_cases a ∈ s <;> simp [*, ← add_assoc, add_right_comm _ 1] theorem card_inter_add_card_union (s t : Finset α) : - (s ∩ t).card + (s ∪ t).card = s.card + t.card := by rw [add_comm, card_union_add_card_inter] + #(s ∩ t) + #(s ∪ t) = #s + #t := by rw [add_comm, card_union_add_card_inter] -lemma card_union (s t : Finset α) : (s ∪ t).card = s.card + t.card - (s ∩ t).card := by +lemma card_union (s t : Finset α) : #(s ∪ t) = #s + #t - #(s ∩ t) := by rw [← card_union_add_card_inter, Nat.add_sub_cancel] -lemma card_inter (s t : Finset α) : (s ∩ t).card = s.card + t.card - (s ∪ t).card := by +lemma card_inter (s t : Finset α) : #(s ∩ t) = #s + #t - #(s ∪ t) := by rw [← card_inter_add_card_union, Nat.add_sub_cancel] -theorem card_union_le (s t : Finset α) : (s ∪ t).card ≤ s.card + t.card := +theorem card_union_le (s t : Finset α) : #(s ∪ t) ≤ #s + #t := card_union_add_card_inter s t ▸ Nat.le_add_right _ _ -lemma card_union_eq_card_add_card : (s ∪ t).card = s.card + t.card ↔ Disjoint s t := by +lemma card_union_eq_card_add_card : #(s ∪ t) = #s + #t ↔ Disjoint s t := by rw [← card_union_add_card_inter]; simp [disjoint_iff_inter_eq_empty] @[simp] alias ⟨_, card_union_of_disjoint⟩ := card_union_eq_card_add_card @@ -489,70 +495,70 @@ lemma card_union_eq_card_add_card : (s ∪ t).card = s.card + t.card ↔ Disjoin @[deprecated (since := "2024-02-09")] alias card_disjoint_union := card_union_of_disjoint lemma cast_card_inter [AddGroupWithOne R] : - ((s ∩ t).card : R) = s.card + t.card - (s ∪ t).card := by + (#(s ∩ t) : R) = #s + #t - #(s ∪ t) := by rw [eq_sub_iff_add_eq, ← cast_add, card_inter_add_card_union, cast_add] lemma cast_card_union [AddGroupWithOne R] : - ((s ∪ t).card : R) = s.card + t.card - (s ∩ t).card := by + (#(s ∪ t) : R) = #s + #t - #(s ∩ t) := by rw [eq_sub_iff_add_eq, ← cast_add, card_union_add_card_inter, cast_add] -theorem card_sdiff (h : s ⊆ t) : card (t \ s) = t.card - s.card := by - suffices card (t \ s) = card (t \ s ∪ s) - s.card by rwa [sdiff_union_of_subset h] at this +theorem card_sdiff (h : s ⊆ t) : #(t \ s) = #t - #s := by + suffices #(t \ s) = #(t \ s ∪ s) - #s by rwa [sdiff_union_of_subset h] at this rw [card_union_of_disjoint sdiff_disjoint, Nat.add_sub_cancel_right] -lemma cast_card_sdiff [AddGroupWithOne R] (h : s ⊆ t) : ((t \ s).card : R) = t.card - s.card := by +lemma cast_card_sdiff [AddGroupWithOne R] (h : s ⊆ t) : (#(t \ s) : R) = #t - #s := by rw [card_sdiff h, Nat.cast_sub (card_mono h)] -theorem card_sdiff_add_card_eq_card {s t : Finset α} (h : s ⊆ t) : card (t \ s) + card s = card t := +theorem card_sdiff_add_card_eq_card {s t : Finset α} (h : s ⊆ t) : #(t \ s) + #s = #t := ((Nat.sub_eq_iff_eq_add (card_le_card h)).mp (card_sdiff h).symm).symm -theorem le_card_sdiff (s t : Finset α) : t.card - s.card ≤ card (t \ s) := +theorem le_card_sdiff (s t : Finset α) : #t - #s ≤ #(t \ s) := calc - card t - card s ≤ card t - card (s ∩ t) := + #t - #s ≤ #t - #(s ∩ t) := Nat.sub_le_sub_left (card_le_card inter_subset_left) _ - _ = card (t \ (s ∩ t)) := (card_sdiff inter_subset_right).symm - _ ≤ card (t \ s) := by rw [sdiff_inter_self_right t s] + _ = #(t \ (s ∩ t)) := (card_sdiff inter_subset_right).symm + _ ≤ #(t \ s) := by rw [sdiff_inter_self_right t s] -theorem card_le_card_sdiff_add_card : s.card ≤ (s \ t).card + t.card := +theorem card_le_card_sdiff_add_card : #s ≤ #(s \ t) + #t := Nat.sub_le_iff_le_add.1 <| le_card_sdiff _ _ -theorem card_sdiff_add_card (s t : Finset α) : (s \ t).card + t.card = (s ∪ t).card := by +theorem card_sdiff_add_card (s t : Finset α) : #(s \ t) + #t = #(s ∪ t) := by rw [← card_union_of_disjoint sdiff_disjoint, sdiff_union_self_eq_union] -lemma card_sdiff_comm (h : s.card = t.card) : (s \ t).card = (t \ s).card := - add_left_injective t.card <| by +lemma card_sdiff_comm (h : #s = #t) : #(s \ t) = #(t \ s) := + add_left_injective #t <| by simp_rw [card_sdiff_add_card, ← h, card_sdiff_add_card, union_comm] @[simp] lemma card_sdiff_add_card_inter (s t : Finset α) : - (s \ t).card + (s ∩ t).card = s.card := by + #(s \ t) + #(s ∩ t) = #s := by rw [← card_union_of_disjoint (disjoint_sdiff_inter _ _), sdiff_union_inter] @[simp] lemma card_inter_add_card_sdiff (s t : Finset α) : - (s ∩ t).card + (s \ t).card = s.card := by + #(s ∩ t) + #(s \ t) = #s := by rw [add_comm, card_sdiff_add_card_inter] /-- **Pigeonhole principle** for two finsets inside an ambient finset. -/ theorem inter_nonempty_of_card_lt_card_add_card (hts : t ⊆ s) (hus : u ⊆ s) - (hstu : s.card < t.card + u.card) : (t ∩ u).Nonempty := by + (hstu : #s < #t + #u) : (t ∩ u).Nonempty := by contrapose! hstu calc - _ = (t ∪ u).card := by simp [← card_union_add_card_inter, not_nonempty_iff_eq_empty.1 hstu] - _ ≤ s.card := by gcongr; exact union_subset hts hus + _ = #(t ∪ u) := by simp [← card_union_add_card_inter, not_nonempty_iff_eq_empty.1 hstu] + _ ≤ #s := by gcongr; exact union_subset hts hus end Lattice theorem filter_card_add_filter_neg_card_eq_card (p : α → Prop) [DecidablePred p] [∀ x, Decidable (¬p x)] : - (s.filter p).card + (s.filter (fun a => ¬ p a)).card = s.card := by + #(s.filter p) + #(s.filter fun a ↦ ¬ p a) = #s := by classical rw [← card_union_of_disjoint (disjoint_filter_filter_neg _ _ _), filter_union_filter_neg_eq] /-- Given a subset `s` of a set `t`, of sizes at most and at least `n` respectively, there exists a set `u` of size `n` which is both a superset of `s` and a subset of `t`. -/ -lemma exists_subsuperset_card_eq (hst : s ⊆ t) (hsn : s.card ≤ n) (hnt : n ≤ t.card) : - ∃ u, s ⊆ u ∧ u ⊆ t ∧ card u = n := by +lemma exists_subsuperset_card_eq (hst : s ⊆ t) (hsn : #s ≤ n) (hnt : n ≤ #t) : + ∃ u, s ⊆ u ∧ u ⊆ t ∧ #u = n := by classical refine Nat.decreasingInduction' ?_ hnt ⟨t, by simp [hst]⟩ intro k _ hnk ⟨u, hu₁, hu₂, hu₃⟩ @@ -561,48 +567,48 @@ lemma exists_subsuperset_card_eq (hst : s ⊆ t) (hsn : s.card ≤ n) (hnt : n exact ⟨u.erase a, by simp [subset_erase, erase_subset_iff_of_mem (hu₂ _), *]⟩ /-- We can shrink a set to any smaller size. -/ -lemma exists_subset_card_eq (hns : n ≤ s.card) : ∃ t ⊆ s, t.card = n := by +lemma exists_subset_card_eq (hns : n ≤ #s) : ∃ t ⊆ s, #t = n := by simpa using exists_subsuperset_card_eq s.empty_subset (by simp) hns /-- Given a set `A` and a set `B` inside it, we can shrink `A` to any appropriate size, and keep `B` inside it. -/ @[deprecated exists_subsuperset_card_eq (since := "2024-06-23")] -theorem exists_intermediate_set {A B : Finset α} (i : ℕ) (h₁ : i + card B ≤ card A) (h₂ : B ⊆ A) : - ∃ C : Finset α, B ⊆ C ∧ C ⊆ A ∧ card C = i + card B := +theorem exists_intermediate_set {A B : Finset α} (i : ℕ) (h₁ : i + #B ≤ #A) (h₂ : B ⊆ A) : + ∃ C : Finset α, B ⊆ C ∧ C ⊆ A ∧ #C = i + #B := exists_subsuperset_card_eq h₂ (Nat.le_add_left ..) h₁ /-- We can shrink `A` to any smaller size. -/ @[deprecated exists_subset_card_eq (since := "2024-06-23")] -theorem exists_smaller_set (A : Finset α) (i : ℕ) (h₁ : i ≤ card A) : - ∃ B : Finset α, B ⊆ A ∧ card B = i := exists_subset_card_eq h₁ +theorem exists_smaller_set (A : Finset α) (i : ℕ) (h₁ : i ≤ #A) : + ∃ B : Finset α, B ⊆ A ∧ #B = i := exists_subset_card_eq h₁ -theorem le_card_iff_exists_subset_card : n ≤ s.card ↔ ∃ t ⊆ s, t.card = n := by +theorem le_card_iff_exists_subset_card : n ≤ #s ↔ ∃ t ⊆ s, #t = n := by refine ⟨fun h => ?_, fun ⟨t, hst, ht⟩ => ht ▸ card_le_card hst⟩ exact exists_subset_card_eq h theorem exists_subset_or_subset_of_two_mul_lt_card [DecidableEq α] {X Y : Finset α} {n : ℕ} - (hXY : 2 * n < (X ∪ Y).card) : ∃ C : Finset α, n < C.card ∧ (C ⊆ X ∨ C ⊆ Y) := by - have h₁ : (X ∩ (Y \ X)).card = 0 := Finset.card_eq_zero.mpr (Finset.inter_sdiff_self X Y) - have h₂ : (X ∪ Y).card = X.card + (Y \ X).card := by + (hXY : 2 * n < #(X ∪ Y)) : ∃ C : Finset α, n < #C ∧ (C ⊆ X ∨ C ⊆ Y) := by + have h₁ : #(X ∩ (Y \ X)) = 0 := Finset.card_eq_zero.mpr (Finset.inter_sdiff_self X Y) + have h₂ : #(X ∪ Y) = #X + #(Y \ X) := by rw [← card_union_add_card_inter X (Y \ X), Finset.union_sdiff_self_eq_union, h₁, add_zero] rw [h₂, Nat.two_mul] at hXY - obtain h | h : n < X.card ∨ n < (Y \ X).card := by contrapose! hXY; omega + obtain h | h : n < #X ∨ n < #(Y \ X) := by contrapose! hXY; omega · exact ⟨X, h, Or.inl (Finset.Subset.refl X)⟩ · exact ⟨Y \ X, h, Or.inr sdiff_subset⟩ /-! ### Explicit description of a finset from its card -/ -theorem card_eq_one : s.card = 1 ↔ ∃ a, s = {a} := by +theorem card_eq_one : #s = 1 ↔ ∃ a, s = {a} := by cases s simp only [Multiset.card_eq_one, Finset.card, ← val_inj, singleton_val] theorem _root_.Multiset.toFinset_card_eq_one_iff [DecidableEq α] (s : Multiset α) : - s.toFinset.card = 1 ↔ Multiset.card s ≠ 0 ∧ ∃ a : α, s = Multiset.card s • {a} := by + #s.toFinset = 1 ↔ Multiset.card s ≠ 0 ∧ ∃ a : α, s = Multiset.card s • {a} := by simp_rw [card_eq_one, Multiset.toFinset_eq_singleton_iff, exists_and_left] theorem exists_eq_insert_iff [DecidableEq α] {s t : Finset α} : - (∃ a ∉ s, insert a s = t) ↔ s ⊆ t ∧ s.card + 1 = t.card := by + (∃ a ∉ s, insert a s = t) ↔ s ⊆ t ∧ #s + 1 = #t := by constructor · rintro ⟨a, ha, rfl⟩ exact ⟨subset_insert _ _, (card_insert_of_not_mem ha).symm⟩ @@ -615,7 +621,7 @@ theorem exists_eq_insert_iff [DecidableEq α] {s t : Finset α} : rw [← ha] exact not_mem_sdiff_of_mem_right hs -theorem card_le_one : s.card ≤ 1 ↔ ∀ a ∈ s, ∀ b ∈ s, a = b := by +theorem card_le_one : #s ≤ 1 ↔ ∀ a ∈ s, ∀ b ∈ s, a = b := by obtain rfl | ⟨x, hx⟩ := s.eq_empty_or_nonempty · simp refine (Nat.succ_le_of_lt (card_pos.2 ⟨x, hx⟩)).le_iff_eq.trans (card_eq_one.trans ⟨?_, ?_⟩) @@ -623,14 +629,14 @@ theorem card_le_one : s.card ≤ 1 ↔ ∀ a ∈ s, ∀ b ∈ s, a = b := by simp · exact fun h => ⟨x, eq_singleton_iff_unique_mem.2 ⟨hx, fun y hy => h _ hy _ hx⟩⟩ -theorem card_le_one_iff : s.card ≤ 1 ↔ ∀ {a b}, a ∈ s → b ∈ s → a = b := by +theorem card_le_one_iff : #s ≤ 1 ↔ ∀ {a b}, a ∈ s → b ∈ s → a = b := by rw [card_le_one] tauto -theorem card_le_one_iff_subsingleton_coe : s.card ≤ 1 ↔ Subsingleton (s : Type _) := +theorem card_le_one_iff_subsingleton_coe : #s ≤ 1 ↔ Subsingleton (s : Type _) := card_le_one.trans (s : Set α).subsingleton_coe.symm -theorem card_le_one_iff_subset_singleton [Nonempty α] : s.card ≤ 1 ↔ ∃ x : α, s ⊆ {x} := by +theorem card_le_one_iff_subset_singleton [Nonempty α] : #s ≤ 1 ↔ ∃ x : α, s ⊆ {x} := by refine ⟨fun H => ?_, ?_⟩ · obtain rfl | ⟨x, hx⟩ := s.eq_empty_or_nonempty · exact ⟨Classical.arbitrary α, empty_subset _⟩ @@ -639,32 +645,32 @@ theorem card_le_one_iff_subset_singleton [Nonempty α] : s.card ≤ 1 ↔ ∃ x rw [← card_singleton x] exact card_le_card hx -lemma exists_mem_ne (hs : 1 < s.card) (a : α) : ∃ b ∈ s, b ≠ a := by +lemma exists_mem_ne (hs : 1 < #s) (a : α) : ∃ b ∈ s, b ≠ a := by have : Nonempty α := ⟨a⟩ by_contra! exact hs.not_le (card_le_one_iff_subset_singleton.2 ⟨a, subset_singleton_iff'.2 this⟩) /-- A `Finset` of a subsingleton type has cardinality at most one. -/ -theorem card_le_one_of_subsingleton [Subsingleton α] (s : Finset α) : s.card ≤ 1 := +theorem card_le_one_of_subsingleton [Subsingleton α] (s : Finset α) : #s ≤ 1 := Finset.card_le_one_iff.2 fun {_ _ _ _} => Subsingleton.elim _ _ -theorem one_lt_card : 1 < s.card ↔ ∃ a ∈ s, ∃ b ∈ s, a ≠ b := by +theorem one_lt_card : 1 < #s ↔ ∃ a ∈ s, ∃ b ∈ s, a ≠ b := by rw [← not_iff_not] push_neg exact card_le_one -theorem one_lt_card_iff : 1 < s.card ↔ ∃ a b, a ∈ s ∧ b ∈ s ∧ a ≠ b := by +theorem one_lt_card_iff : 1 < #s ↔ ∃ a b, a ∈ s ∧ b ∈ s ∧ a ≠ b := by rw [one_lt_card] simp only [exists_prop, exists_and_left] -theorem one_lt_card_iff_nontrivial : 1 < s.card ↔ s.Nontrivial := by +theorem one_lt_card_iff_nontrivial : 1 < #s ↔ s.Nontrivial := by rw [← not_iff_not, not_lt, Finset.Nontrivial, ← Set.nontrivial_coe_sort, not_nontrivial_iff_subsingleton, card_le_one_iff_subsingleton_coe, coe_sort_coe] @[deprecated (since := "2024-02-05")] alias one_lt_card_iff_nontrivial_coe := one_lt_card_iff_nontrivial -theorem exists_ne_of_one_lt_card (hs : 1 < s.card) (a : α) : ∃ b, b ∈ s ∧ b ≠ a := by +theorem exists_ne_of_one_lt_card (hs : 1 < #s) (a : α) : ∃ b, b ∈ s ∧ b ≠ a := by obtain ⟨x, hx, y, hy, hxy⟩ := Finset.one_lt_card.mp hs by_cases ha : y = a · exact ⟨x, hx, ne_of_ne_of_eq hxy ha⟩ @@ -674,8 +680,8 @@ theorem exists_ne_of_one_lt_card (hs : 1 < s.card) (a : α) : ∃ b, b ∈ s ∧ its projection to some factor is nontrivial, and the fibers of the projection are proper subsets. -/ lemma exists_of_one_lt_card_pi {ι : Type*} {α : ι → Type*} [∀ i, DecidableEq (α i)] - {s : Finset (∀ i, α i)} (h : 1 < s.card) : - ∃ i, 1 < (s.image (· i)).card ∧ ∀ ai, s.filter (· i = ai) ⊂ s := by + {s : Finset (∀ i, α i)} (h : 1 < #s) : + ∃ i, 1 < #(s.image (· i)) ∧ ∀ ai, s.filter (· i = ai) ⊂ s := by simp_rw [one_lt_card_iff, Function.ne_iff] at h ⊢ obtain ⟨a1, a2, h1, h2, i, hne⟩ := h refine ⟨i, ⟨_, _, mem_image_of_mem _ h1, mem_image_of_mem _ h2, hne⟩, fun ai => ?_⟩ @@ -684,21 +690,21 @@ lemma exists_of_one_lt_card_pi {ι : Type*} {α : ι → Type*} [∀ i, Decidabl exacts [⟨a1, h1, hne⟩, ⟨a2, h2, hne⟩] theorem card_eq_succ_iff_cons : - s.card = n + 1 ↔ ∃ a t, ∃ (h : a ∉ t), cons a t h = s ∧ t.card = n := + #s = n + 1 ↔ ∃ a t, ∃ (h : a ∉ t), cons a t h = s ∧ #t = n := ⟨cons_induction_on s (by simp) fun a s _ _ _ => ⟨a, s, by simp_all⟩, fun ⟨a, t, _, hs, _⟩ => by simpa [← hs]⟩ section DecidableEq variable [DecidableEq α] -theorem card_eq_succ : s.card = n + 1 ↔ ∃ a t, a ∉ t ∧ insert a t = s ∧ t.card = n := +theorem card_eq_succ : #s = n + 1 ↔ ∃ a t, a ∉ t ∧ insert a t = s ∧ #t = n := ⟨fun h => - let ⟨a, has⟩ := card_pos.mp (h.symm ▸ Nat.zero_lt_succ _ : 0 < s.card) + let ⟨a, has⟩ := card_pos.mp (h.symm ▸ Nat.zero_lt_succ _ : 0 < #s) ⟨a, s.erase a, s.not_mem_erase a, insert_erase has, by simp only [h, card_erase_of_mem has, Nat.add_sub_cancel_right]⟩, fun ⟨_, _, hat, s_eq, n_eq⟩ => s_eq ▸ n_eq ▸ card_insert_of_not_mem hat⟩ -theorem card_eq_two : s.card = 2 ↔ ∃ x y, x ≠ y ∧ s = {x, y} := by +theorem card_eq_two : #s = 2 ↔ ∃ x y, x ≠ y ∧ s = {x, y} := by constructor · rw [card_eq_succ] simp_rw [card_eq_one] @@ -707,7 +713,7 @@ theorem card_eq_two : s.card = 2 ↔ ∃ x y, x ≠ y ∧ s = {x, y} := by · rintro ⟨x, y, h, rfl⟩ exact card_pair h -theorem card_eq_three : s.card = 3 ↔ ∃ x y z, x ≠ y ∧ x ≠ z ∧ y ≠ z ∧ s = {x, y, z} := by +theorem card_eq_three : #s = 3 ↔ ∃ x y z, x ≠ y ∧ x ≠ z ∧ y ≠ z ∧ s = {x, y, z} := by constructor · rw [card_eq_succ] simp_rw [card_eq_two] @@ -720,7 +726,7 @@ theorem card_eq_three : s.card = 3 ↔ ∃ x y z, x ≠ y ∧ x ≠ z ∧ y ≠ end DecidableEq -theorem two_lt_card_iff : 2 < s.card ↔ ∃ a b c, a ∈ s ∧ b ∈ s ∧ c ∈ s ∧ a ≠ b ∧ a ≠ c ∧ b ≠ c := by +theorem two_lt_card_iff : 2 < #s ↔ ∃ a b c, a ∈ s ∧ b ∈ s ∧ c ∈ s ∧ a ≠ b ∧ a ≠ c ∧ b ≠ c := by classical simp_rw [lt_iff_add_one_le, le_card_iff_exists_subset_card, reduceAdd, card_eq_three, ← exists_and_left, exists_comm (α := Finset α)] @@ -730,7 +736,7 @@ theorem two_lt_card_iff : 2 < s.card ↔ ∃ a b c, a ∈ s ∧ b ∈ s ∧ c · rintro ⟨a, b, c, ha, hb, hc, hab, hac, hbc⟩ exact ⟨a, b, c, {a, b, c}, by simp_all [insert_subset_iff]⟩ -theorem two_lt_card : 2 < s.card ↔ ∃ a ∈ s, ∃ b ∈ s, ∃ c ∈ s, a ≠ b ∧ a ≠ c ∧ b ≠ c := by +theorem two_lt_card : 2 < #s ↔ ∃ a ∈ s, ∃ b ∈ s, ∃ c ∈ s, a ≠ b ∧ a ≠ c ∧ b ≠ c := by simp_rw [two_lt_card_iff, exists_and_left] /-! ### Inductions -/ @@ -743,9 +749,9 @@ def strongInduction {p : Finset α → Sort*} (H : ∀ s, (∀ t ⊂ s, p t) → ∀ s : Finset α, p s | s => H s fun t h => - have : t.card < s.card := card_lt_card h + have : #t < #s := card_lt_card h strongInduction H t - termination_by s => Finset.card s + termination_by s => #s @[nolint unusedHavesSuffices] -- Porting note: false positive theorem strongInduction_eq {p : Finset α → Sort*} (H : ∀ s, (∀ t ⊂ s, p t) → p s) @@ -789,25 +795,25 @@ protected lemma Nonempty.strong_induction {p : ∀ s, s.Nonempty → Prop} · refine h₁ hs fun t ht hts ↦ ?_ have := card_lt_card hts exact ht.strong_induction h₀ h₁ -termination_by s => Finset.card s +termination_by s => #s /-- Suppose that, given that `p t` can be defined on all supersets of `s` of cardinality less than `n`, one knows how to define `p s`. Then one can inductively define `p s` for all finsets `s` of cardinality less than `n`, starting from finsets of card `n` and iterating. This can be used either to define data, or to prove properties. -/ def strongDownwardInduction {p : Finset α → Sort*} {n : ℕ} - (H : ∀ t₁, (∀ {t₂ : Finset α}, t₂.card ≤ n → t₁ ⊂ t₂ → p t₂) → t₁.card ≤ n → p t₁) : - ∀ s : Finset α, s.card ≤ n → p s + (H : ∀ t₁, (∀ {t₂ : Finset α}, #t₂ ≤ n → t₁ ⊂ t₂ → p t₂) → #t₁ ≤ n → p t₁) : + ∀ s : Finset α, #s ≤ n → p s | s => H s fun {t} ht h => have := Finset.card_lt_card h - have : n - t.card < n - s.card := by omega + have : n - #t < n - #s := by omega strongDownwardInduction H t ht - termination_by s => n - s.card + termination_by s => n - #s @[nolint unusedHavesSuffices] -- Porting note: false positive theorem strongDownwardInduction_eq {p : Finset α → Sort*} - (H : ∀ t₁, (∀ {t₂ : Finset α}, t₂.card ≤ n → t₁ ⊂ t₂ → p t₂) → t₁.card ≤ n → p t₁) + (H : ∀ t₁, (∀ {t₂ : Finset α}, #t₂ ≤ n → t₁ ⊂ t₂ → p t₂) → #t₁ ≤ n → p t₁) (s : Finset α) : strongDownwardInduction H s = H s fun {t} ht _ => strongDownwardInduction H t ht := by rw [strongDownwardInduction] @@ -815,13 +821,13 @@ theorem strongDownwardInduction_eq {p : Finset α → Sort*} /-- Analogue of `strongDownwardInduction` with order of arguments swapped. -/ @[elab_as_elim] def strongDownwardInductionOn {p : Finset α → Sort*} (s : Finset α) - (H : ∀ t₁, (∀ {t₂ : Finset α}, t₂.card ≤ n → t₁ ⊂ t₂ → p t₂) → t₁.card ≤ n → p t₁) : - s.card ≤ n → p s := + (H : ∀ t₁, (∀ {t₂ : Finset α}, #t₂ ≤ n → t₁ ⊂ t₂ → p t₂) → #t₁ ≤ n → p t₁) : + #s ≤ n → p s := strongDownwardInduction H s @[nolint unusedHavesSuffices] -- Porting note: false positive theorem strongDownwardInductionOn_eq {p : Finset α → Sort*} (s : Finset α) - (H : ∀ t₁, (∀ {t₂ : Finset α}, t₂.card ≤ n → t₁ ⊂ t₂ → p t₂) → t₁.card ≤ n → p t₁) : + (H : ∀ t₁, (∀ {t₂ : Finset α}, #t₂ ≤ n → t₁ ⊂ t₂ → p t₂) → #t₁ ≤ n → p t₁) : s.strongDownwardInductionOn H = H s fun {t} ht _ => t.strongDownwardInductionOn H ht := by dsimp only [strongDownwardInductionOn] rw [strongDownwardInduction] From e007f8f7404b132837d23260b6733539080a0511 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 17 Oct 2024 15:41:53 +0000 Subject: [PATCH 102/505] =?UTF-8?q?feat(LinearAlgebra/Quotient):=20`(?= =?UTF-8?q?=E2=88=80=20a=20:=20=20M=20=E2=A7=B8=20p,=20P=20a)=20=E2=86=94?= =?UTF-8?q?=20=E2=88=80=20a=20:=20M,=20P=20=E2=9F=A6a=E2=9F=A7`=20(#17808)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit From PFR --- Mathlib/LinearAlgebra/Quotient.lean | 2 ++ 1 file changed, 2 insertions(+) diff --git a/Mathlib/LinearAlgebra/Quotient.lean b/Mathlib/LinearAlgebra/Quotient.lean index 9f3e755dd20eb..ca687f953d024 100644 --- a/Mathlib/LinearAlgebra/Quotient.lean +++ b/Mathlib/LinearAlgebra/Quotient.lean @@ -104,6 +104,8 @@ theorem mk_neg : (mk (-x) : M ⧸ p) = -(mk x : M ⧸ p) := theorem mk_sub : (mk (x - y) : M ⧸ p) = (mk x : M ⧸ p) - (mk y : M ⧸ p) := rfl +protected nonrec lemma «forall» {P : M ⧸ p → Prop} : (∀ a, P a) ↔ ∀ a, P (mk a) := Quotient.forall + section SMul variable {S : Type*} [SMul S R] [SMul S M] [IsScalarTower S R M] (P : Submodule R M) From d2ce9970d101778fc9d13ea498ab229461ed3c07 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Thu, 17 Oct 2024 15:41:55 +0000 Subject: [PATCH 103/505] chore(Analysis): remove unused variables (#17874) These were already unused on the master branch but became exposed to the linter after #17870 changed some `CoeFun` instances. --- Mathlib/Analysis/Analytic/Basic.lean | 2 +- Mathlib/Analysis/Analytic/Inverse.lean | 6 +++--- Mathlib/Analysis/Calculus/IteratedDeriv/Defs.lean | 2 +- Mathlib/Analysis/Normed/Algebra/TrivSqZeroExt.lean | 2 +- 4 files changed, 6 insertions(+), 6 deletions(-) diff --git a/Mathlib/Analysis/Analytic/Basic.lean b/Mathlib/Analysis/Analytic/Basic.lean index 04e001a4a25bc..35af733c58d75 100644 --- a/Mathlib/Analysis/Analytic/Basic.lean +++ b/Mathlib/Analysis/Analytic/Basic.lean @@ -638,7 +638,7 @@ theorem HasFPowerSeriesWithinOnBall.coeff_zero (hf : HasFPowerSeriesWithinOnBall (v : Fin 0 → E) : pf 0 v = f x := by have v_eq : v = fun i => 0 := Subsingleton.elim _ _ have zero_mem : (0 : E) ∈ EMetric.ball (0 : E) r := by simp [hf.r_pos] - have : ∀ i, i ≠ 0 → (pf i fun j => 0) = 0 := by + have : ∀ i, i ≠ 0 → (pf i fun _ => 0) = 0 := by intro i hi have : 0 < i := pos_iff_ne_zero.2 hi exact ContinuousMultilinearMap.map_coord_zero _ (⟨0, this⟩ : Fin i) rfl diff --git a/Mathlib/Analysis/Analytic/Inverse.lean b/Mathlib/Analysis/Analytic/Inverse.lean index eadc577c91124..34ad1130c225c 100644 --- a/Mathlib/Analysis/Analytic/Inverse.lean +++ b/Mathlib/Analysis/Analytic/Inverse.lean @@ -272,7 +272,7 @@ theorem rightInv_coeff (p : FormalMultilinearSeries 𝕜 E F) (i : E ≃L[𝕜] congr (config := { closePost := false }) 1 ext v have N : 0 < n + 2 := by norm_num - have : ((p 1) fun i : Fin 1 => 0) = 0 := ContinuousMultilinearMap.map_zero _ + have : ((p 1) fun _ : Fin 1 => 0) = 0 := ContinuousMultilinearMap.map_zero _ simp [comp_rightInv_aux1 N, lt_irrefl n, this, comp_rightInv_aux2, -Set.toFinset_setOf] /-! ### Coincidence of the left and the right inverse -/ @@ -630,10 +630,10 @@ lemma HasFPowerSeriesAt.eventually_hasSum_of_comp {f : E → F} {g : F → G} simp only [id_eq, eventually_atTop, ge_iff_le] rcases mem_nhds_iff.1 hu with ⟨v, vu, v_open, hv⟩ obtain ⟨a₀, b₀, hab⟩ : ∃ a₀ b₀, ∀ (a b : ℕ), a₀ ≤ a → b₀ ≤ b → - q.partialSum a (p.partialSum b y - (p 0) fun x ↦ 0) ∈ v := by + q.partialSum a (p.partialSum b y - (p 0) fun _ ↦ 0) ∈ v := by simpa using hy (v_open.mem_nhds hv) refine ⟨a₀, fun a ha ↦ ?_⟩ - have : Tendsto (fun b ↦ q.partialSum a (p.partialSum b y - (p 0) fun x ↦ 0)) atTop + have : Tendsto (fun b ↦ q.partialSum a (p.partialSum b y - (p 0) fun _ ↦ 0)) atTop (𝓝 (q.partialSum a (f (x + y) - f x))) := by have : ContinuousAt (q.partialSum a) (f (x + y) - f x) := (partialSum_continuous q a).continuousAt diff --git a/Mathlib/Analysis/Calculus/IteratedDeriv/Defs.lean b/Mathlib/Analysis/Calculus/IteratedDeriv/Defs.lean index f5790c73a631f..cf2ae12780d4c 100644 --- a/Mathlib/Analysis/Calculus/IteratedDeriv/Defs.lean +++ b/Mathlib/Analysis/Calculus/IteratedDeriv/Defs.lean @@ -166,7 +166,7 @@ theorem iteratedDerivWithin_succ {x : 𝕜} (hxs : UniqueDiffWithinAt 𝕜 s x) rw [iteratedDerivWithin_eq_iteratedFDerivWithin, iteratedFDerivWithin_succ_apply_left, iteratedFDerivWithin_eq_equiv_comp, LinearIsometryEquiv.comp_fderivWithin _ hxs, derivWithin] change ((ContinuousMultilinearMap.mkPiRing 𝕜 (Fin n) ((fderivWithin 𝕜 - (iteratedDerivWithin n f s) s x : 𝕜 → F) 1) : (Fin n → 𝕜) → F) fun i : Fin n => 1) = + (iteratedDerivWithin n f s) s x : 𝕜 → F) 1) : (Fin n → 𝕜) → F) fun _ : Fin n => 1) = (fderivWithin 𝕜 (iteratedDerivWithin n f s) s x : 𝕜 → F) 1 simp diff --git a/Mathlib/Analysis/Normed/Algebra/TrivSqZeroExt.lean b/Mathlib/Analysis/Normed/Algebra/TrivSqZeroExt.lean index 64ad667b3190a..c3e4233f09ac9 100644 --- a/Mathlib/Analysis/Normed/Algebra/TrivSqZeroExt.lean +++ b/Mathlib/Analysis/Normed/Algebra/TrivSqZeroExt.lean @@ -115,7 +115,7 @@ variable [T2Space R] [T2Space M] theorem exp_def_of_smul_comm (x : tsze R M) (hx : MulOpposite.op x.fst • x.snd = x.fst • x.snd) : exp 𝕜 x = inl (exp 𝕜 x.fst) + inr (exp 𝕜 x.fst • x.snd) := by simp_rw [exp, FormalMultilinearSeries.sum] - by_cases h : Summable (fun (n : ℕ) => (expSeries 𝕜 R n) fun x_1 ↦ fst x) + by_cases h : Summable (fun (n : ℕ) => (expSeries 𝕜 R n) fun _ ↦ fst x) · refine (hasSum_expSeries_of_smul_comm 𝕜 x hx ?_).tsum_eq exact h.hasSum · rw [tsum_eq_zero_of_not_summable h, zero_smul, inr_zero, inl_zero, zero_add, From d3038070bcab78ca09317951523576e264e9cf84 Mon Sep 17 00:00:00 2001 From: damiano Date: Thu, 17 Oct 2024 16:19:08 +0000 Subject: [PATCH 104/505] chore: another removal of unused variables (#17790) --- Mathlib/Algebra/Field/ULift.lean | 4 +-- Mathlib/Algebra/Group/Action/Basic.lean | 2 +- Mathlib/Algebra/Group/Indicator.lean | 6 ++-- Mathlib/Algebra/Group/Pi/Lemmas.lean | 3 +- Mathlib/Algebra/Group/Submonoid/Defs.lean | 4 --- Mathlib/Algebra/Group/Subsemigroup/Defs.lean | 4 +-- .../Algebra/GroupWithZero/Action/Basic.lean | 2 +- Mathlib/Algebra/GroupWithZero/Action/Pi.lean | 5 +--- Mathlib/Algebra/Module/Defs.lean | 10 +++---- Mathlib/Algebra/Module/Equiv/Defs.lean | 12 +++----- Mathlib/Algebra/Module/LinearMap/Defs.lean | 28 ++++++++----------- Mathlib/Algebra/Module/Pi.lean | 3 -- Mathlib/Algebra/Module/Submodule/Range.lean | 19 ++++++------- Mathlib/Algebra/Order/Interval/Set/Group.lean | 2 +- Mathlib/Algebra/Order/Ring/Abs.lean | 8 +++--- Mathlib/Algebra/Regular/SMul.lean | 3 +- Mathlib/Algebra/Ring/Pi.lean | 3 +- Mathlib/Algebra/Ring/Prod.lean | 2 +- .../Analysis/InnerProductSpace/Symmetric.lean | 10 ++----- .../OperatorNorm/Completeness.lean | 11 ++++---- Mathlib/Data/Fin/VecNotation.lean | 2 +- Mathlib/Data/Int/Cast/Lemmas.lean | 6 ++-- Mathlib/Data/Int/Lemmas.lean | 2 -- Mathlib/Data/Int/Order/Lemmas.lean | 3 -- Mathlib/Data/List/Perm.lean | 2 +- Mathlib/Data/List/Sort.lean | 2 +- Mathlib/Data/Nat/Cast/Order/Ring.lean | 2 +- Mathlib/Dynamics/FixedPoints/Basic.lean | 2 +- .../Geometry/Manifold/Algebra/LieGroup.lean | 7 ++--- .../Manifold/VectorBundle/SmoothSection.lean | 10 ++----- Mathlib/GroupTheory/Congruence/Defs.lean | 4 +-- Mathlib/GroupTheory/CosetCover.lean | 2 +- Mathlib/GroupTheory/DoubleCoset.lean | 2 +- Mathlib/GroupTheory/GroupAction/Hom.lean | 5 +--- Mathlib/GroupTheory/NoncommPiCoprod.lean | 9 ------ Mathlib/GroupTheory/Sylow.lean | 4 +-- Mathlib/LinearAlgebra/Basis/Basic.lean | 15 ++++------ .../LinearAlgebra/FiniteDimensional/Defs.lean | 9 ++---- Mathlib/MeasureTheory/Integral/Bochner.lean | 21 ++++++-------- Mathlib/Order/Bounds/Basic.lean | 2 +- Mathlib/Order/CompleteLattice.lean | 2 +- Mathlib/Order/GaloisConnection.lean | 6 ++-- Mathlib/Order/PiLex.lean | 2 +- .../RingTheory/Localization/Submodule.lean | 9 ++---- Mathlib/Topology/Algebra/Module/Basic.lean | 14 ++++------ 45 files changed, 106 insertions(+), 179 deletions(-) diff --git a/Mathlib/Algebra/Field/ULift.lean b/Mathlib/Algebra/Field/ULift.lean index ccb6baccbc575..215b218b67831 100644 --- a/Mathlib/Algebra/Field/ULift.lean +++ b/Mathlib/Algebra/Field/ULift.lean @@ -15,8 +15,8 @@ This file defines instances for field, semifield and related structures on `ULif (Recall `ULift α` is just a "copy" of a type `α` in a higher universe.) -/ -universe u v -variable {α : Type u} {x y : ULift.{v} α} +universe u +variable {α : Type u} namespace ULift diff --git a/Mathlib/Algebra/Group/Action/Basic.lean b/Mathlib/Algebra/Group/Action/Basic.lean index 8515f84bac94f..0077326ec936b 100644 --- a/Mathlib/Algebra/Group/Action/Basic.lean +++ b/Mathlib/Algebra/Group/Action/Basic.lean @@ -16,7 +16,7 @@ This file contains lemmas about group actions that require more imports than assert_not_exists MonoidWithZero -variable {α β γ : Type*} +variable {α β : Type*} section MulAction diff --git a/Mathlib/Algebra/Group/Indicator.lean b/Mathlib/Algebra/Group/Indicator.lean index fc2d14cdfb140..38b9e8ce30021 100644 --- a/Mathlib/Algebra/Group/Indicator.lean +++ b/Mathlib/Algebra/Group/Indicator.lean @@ -33,7 +33,7 @@ assert_not_exists MonoidWithZero open Function -variable {α β ι M N : Type*} +variable {α β M N : Type*} namespace Set @@ -262,7 +262,7 @@ end One section Monoid -variable [MulOneClass M] {s t : Set α} {f g : α → M} {a : α} +variable [MulOneClass M] {s t : Set α} {a : α} @[to_additive] theorem mulIndicator_union_mul_inter_apply (f : α → M) (s t : Set α) (a : α) : @@ -360,7 +360,7 @@ end Monoid section Group -variable {G : Type*} [Group G] {s t : Set α} {f g : α → G} {a : α} +variable {G : Type*} [Group G] {s t : Set α} @[to_additive] theorem mulIndicator_inv' (s : Set α) (f : α → G) : mulIndicator s f⁻¹ = (mulIndicator s f)⁻¹ := diff --git a/Mathlib/Algebra/Group/Pi/Lemmas.lean b/Mathlib/Algebra/Group/Pi/Lemmas.lean index 1dee66ec01d1b..0f412767df4f9 100644 --- a/Mathlib/Algebra/Group/Pi/Lemmas.lean +++ b/Mathlib/Algebra/Group/Pi/Lemmas.lean @@ -26,8 +26,7 @@ variable {I : Type u} -- The indexing type variable {f : I → Type v} --- The family of types already equipped with instances -variable (x y : ∀ i, f i) (i j : I) +variable (i : I) @[to_additive (attr := simp)] theorem Set.range_one {α β : Type*} [One β] [Nonempty α] : Set.range (1 : α → β) = {1} := diff --git a/Mathlib/Algebra/Group/Submonoid/Defs.lean b/Mathlib/Algebra/Group/Submonoid/Defs.lean index 4f005ea640695..d1a62b639d86a 100644 --- a/Mathlib/Algebra/Group/Submonoid/Defs.lean +++ b/Mathlib/Algebra/Group/Submonoid/Defs.lean @@ -47,15 +47,11 @@ submonoid, submonoids assert_not_exists CompleteLattice assert_not_exists MonoidWithZero --- Only needed for notation --- Only needed for notation variable {M : Type*} {N : Type*} -variable {A : Type*} section NonAssoc variable [MulOneClass M] {s : Set M} -variable [AddZeroClass A] {t : Set A} /-- `OneMemClass S M` says `S` is a type of subsets `s ≤ M`, such that `1 ∈ s` for all `s`. -/ class OneMemClass (S : Type*) (M : outParam Type*) [One M] [SetLike S M] : Prop where diff --git a/Mathlib/Algebra/Group/Subsemigroup/Defs.lean b/Mathlib/Algebra/Group/Subsemigroup/Defs.lean index 946d325924172..aa0008eee16ba 100644 --- a/Mathlib/Algebra/Group/Subsemigroup/Defs.lean +++ b/Mathlib/Algebra/Group/Subsemigroup/Defs.lean @@ -47,13 +47,11 @@ subsemigroup, subsemigroups assert_not_exists CompleteLattice assert_not_exists MonoidWithZero --- Only needed for notation -variable {M : Type*} {N : Type*} {A : Type*} +variable {M : Type*} {N : Type*} section NonAssoc variable [Mul M] {s : Set M} -variable [Add A] {t : Set A} /-- `MulMemClass S M` says `S` is a type of sets `s : Set M` that are closed under `(*)` -/ class MulMemClass (S : Type*) (M : outParam Type*) [Mul M] [SetLike S M] : Prop where diff --git a/Mathlib/Algebra/GroupWithZero/Action/Basic.lean b/Mathlib/Algebra/GroupWithZero/Action/Basic.lean index f61edb220be6b..39b682e9f1e61 100644 --- a/Mathlib/Algebra/GroupWithZero/Action/Basic.lean +++ b/Mathlib/Algebra/GroupWithZero/Action/Basic.lean @@ -68,7 +68,7 @@ protected lemma MulAction.surjective₀ (ha : a ≠ 0) : Surjective (a • · : end GroupWithZero section DistribMulAction -variable [Group G] [Monoid M] [AddMonoid A] [DistribMulAction M A] +variable [Group G] [Monoid M] [AddMonoid A] variable (A) /-- Each element of the group defines an additive monoid isomorphism. diff --git a/Mathlib/Algebra/GroupWithZero/Action/Pi.lean b/Mathlib/Algebra/GroupWithZero/Action/Pi.lean index a68e1a97b1d8a..f465f4b4caa56 100644 --- a/Mathlib/Algebra/GroupWithZero/Action/Pi.lean +++ b/Mathlib/Algebra/GroupWithZero/Action/Pi.lean @@ -21,16 +21,13 @@ This file defines instances for `MulActionWithZero` and related structures on `P -/ -universe u v w +universe u v variable {I : Type u} -- The indexing type variable {f : I → Type v} --- The family of types already equipped with instances -variable (x y : ∀ i, f i) (i : I) - namespace Pi instance smulZeroClass (α) {n : ∀ i, Zero <| f i} [∀ i, SMulZeroClass α <| f i] : diff --git a/Mathlib/Algebra/Module/Defs.lean b/Mathlib/Algebra/Module/Defs.lean index f810b34644f19..a024554de986a 100644 --- a/Mathlib/Algebra/Module/Defs.lean +++ b/Mathlib/Algebra/Module/Defs.lean @@ -46,7 +46,7 @@ open Function Set universe u v -variable {α R k S M M₂ M₃ ι : Type*} +variable {R S M M₂ : Type*} /-- A module is a generalization of vector spaces to a scalar semiring. It consists of a scalar semiring `R` and an additive monoid of "vectors" `M`, @@ -63,7 +63,7 @@ class Module (R : Type u) (M : Type v) [Semiring R] [AddCommMonoid M] extends section AddCommMonoid -variable [Semiring R] [AddCommMonoid M] [Module R M] (r s : R) (x y : M) +variable [Semiring R] [AddCommMonoid M] [Module R M] (r s : R) (x : M) -- see Note [lower instance priority] /-- A module over a semiring automatically inherits a `MulActionWithZero` structure. -/ @@ -218,7 +218,7 @@ theorem Module.ext' {R : Type*} [Semiring R] {M : Type*} [AddCommMonoid M] (P Q section Module -variable [Ring R] [AddCommGroup M] [Module R M] (r s : R) (x y : M) +variable [Ring R] [AddCommGroup M] [Module R M] (r : R) (x : M) @[simp] theorem neg_smul : -r • x = -(r • x) := @@ -360,7 +360,7 @@ end AddCommMonoid section AddCommGroup -variable [Semiring S] [Ring R] [AddCommGroup M] [Module S M] [Module R M] +variable [Ring R] [AddCommGroup M] [Module R M] section @@ -483,7 +483,7 @@ theorem two_nsmul_eq_zero end Nat -variable [Semiring R] [AddCommMonoid M] [Module R M] +variable [Semiring R] variable (R M) /-- If `M` is an `R`-module with one and `M` has characteristic zero, then `R` has characteristic diff --git a/Mathlib/Algebra/Module/Equiv/Defs.lean b/Mathlib/Algebra/Module/Equiv/Defs.lean index 9a5d7dcbdfd0d..7d08836e2d79f 100644 --- a/Mathlib/Algebra/Module/Equiv/Defs.lean +++ b/Mathlib/Algebra/Module/Equiv/Defs.lean @@ -38,11 +38,9 @@ assert_not_exists Pi.module open Function -universe u u' v w x y z - variable {R : Type*} {R₁ : Type*} {R₂ : Type*} {R₃ : Type*} -variable {k : Type*} {K : Type*} {S : Type*} {M : Type*} {M₁ : Type*} {M₂ : Type*} {M₃ : Type*} -variable {N₁ : Type*} {N₂ : Type*} {N₃ : Type*} {N₄ : Type*} {ι : Type*} +variable {S : Type*} {M : Type*} {M₁ : Type*} {M₂ : Type*} {M₃ : Type*} +variable {N₁ : Type*} {N₂ : Type*} section @@ -133,7 +131,6 @@ namespace LinearEquiv section AddCommMonoid -variable {M₄ : Type*} variable [Semiring R] [Semiring S] section @@ -198,7 +195,7 @@ section variable [Semiring R₁] [Semiring R₂] [Semiring R₃] variable [AddCommMonoid M] [AddCommMonoid M₁] [AddCommMonoid M₂] -variable [AddCommMonoid M₃] [AddCommMonoid M₄] +variable [AddCommMonoid M₃] variable [AddCommMonoid N₁] [AddCommMonoid N₂] variable {module_M : Module R M} {module_S_M₂ : Module S M₂} {σ : R →+* S} {σ' : S →+* R} variable {re₁ : RingHomInvPair σ σ'} {re₂ : RingHomInvPair σ' σ} @@ -506,8 +503,7 @@ def _root_.RingEquiv.toSemilinearEquiv (f : R ≃+* S) : toFun := f map_smul' := f.map_mul } -variable [Semiring R₁] [Semiring R₂] [Semiring R₃] -variable [AddCommMonoid M] [AddCommMonoid M₁] [AddCommMonoid M₂] +variable [AddCommMonoid M] /-- An involutive linear map is a linear equivalence. -/ def ofInvolutive {σ σ' : R →+* R} [RingHomInvPair σ σ'] [RingHomInvPair σ' σ] diff --git a/Mathlib/Algebra/Module/LinearMap/Defs.lean b/Mathlib/Algebra/Module/LinearMap/Defs.lean index 0a75f33f5bf9c..21a6e98c5445a 100644 --- a/Mathlib/Algebra/Module/LinearMap/Defs.lean +++ b/Mathlib/Algebra/Module/LinearMap/Defs.lean @@ -57,9 +57,9 @@ assert_not_exists Field open Function -universe u u' v w x y z +universe u u' v w -variable {R R₁ R₂ R₃ k S S₃ T M M₁ M₂ M₃ N₁ N₂ N₃ ι : Type*} +variable {R R₁ R₂ R₃ S S₃ T M M₁ M₂ M₃ N₂ N₃ : Type*} /-- A map `f` between modules over a semiring is linear if it satisfies the two properties `f (x + y) = f x + f y` and `f (c • x) = c • f x`. The predicate `IsLinearMap R f` asserts this @@ -194,7 +194,6 @@ variable [Semiring R] [Semiring S] section variable [AddCommMonoid M] [AddCommMonoid M₁] [AddCommMonoid M₂] [AddCommMonoid M₃] -variable [AddCommMonoid N₁] [AddCommMonoid N₂] [AddCommMonoid N₃] variable [Module R M] [Module R M₂] [Module S M₃] variable {σ : R →+* S} @@ -299,15 +298,14 @@ end section variable [AddCommMonoid M] [AddCommMonoid M₁] [AddCommMonoid M₂] [AddCommMonoid M₃] -variable [AddCommMonoid N₁] [AddCommMonoid N₂] [AddCommMonoid N₃] variable [Module R M] [Module R M₂] [Module S M₃] variable (σ : R →+* S) -variable (fₗ gₗ : M →ₗ[R] M₂) (f g : M →ₛₗ[σ] M₃) +variable (fₗ : M →ₗ[R] M₂) (f g : M →ₛₗ[σ] M₃) theorem isLinear : IsLinearMap R fₗ := ⟨fₗ.map_add', fₗ.map_smul'⟩ -variable {fₗ gₗ f g σ} +variable {fₗ f g σ} theorem coe_injective : Injective (DFunLike.coe : (M →ₛₗ[σ] M₃) → _) := DFunLike.coe_injective @@ -323,7 +321,7 @@ protected theorem congr_fun (h : f = g) (x : M) : f x = g x := theorem mk_coe (f : M →ₛₗ[σ] M₃) (h) : (LinearMap.mk f h : M →ₛₗ[σ] M₃) = f := rfl -variable (fₗ gₗ f g) +variable (fₗ f g) protected theorem map_add (x y : M) : f (x + y) = f x + f y := map_add f x y @@ -537,7 +535,7 @@ theorem cancel_left (hf : Injective f) : f.comp g = f.comp g' ↔ g = g' := end -variable [AddCommMonoid M] [AddCommMonoid M₁] [AddCommMonoid M₂] [AddCommMonoid M₃] +variable [AddCommMonoid M] [AddCommMonoid M₂] [AddCommMonoid M₃] /-- If a function `g` is a left and right inverse of a linear map `f`, then `g` is linear itself. -/ def inverse [Module R M] [Module S M₂] {σ : R →+* S} {σ' : S →+* R} [RingHomInvPair σ σ'] @@ -721,12 +719,11 @@ namespace LinearMap section SMul -variable [Semiring R] [Semiring R₂] [Semiring R₃] -variable [AddCommMonoid M] [AddCommMonoid M₂] [AddCommMonoid M₃] -variable [Module R M] [Module R₂ M₂] [Module R₃ M₃] -variable {σ₁₂ : R →+* R₂} {σ₂₃ : R₂ →+* R₃} {σ₁₃ : R →+* R₃} [RingHomCompTriple σ₁₂ σ₂₃ σ₁₃] +variable [Semiring R] [Semiring R₂] +variable [AddCommMonoid M] [AddCommMonoid M₂] +variable [Module R M] [Module R₂ M₂] +variable {σ₁₂ : R →+* R₂} variable [Monoid S] [DistribMulAction S M₂] [SMulCommClass R₂ S M₂] -variable [Monoid S₃] [DistribMulAction S₃ M₃] [SMulCommClass R₃ S₃ M₃] variable [Monoid T] [DistribMulAction T M₂] [SMulCommClass R₂ T M₂] instance : SMul S (M →ₛₗ[σ₁₂] M₂) := @@ -762,9 +759,9 @@ section Arithmetic variable [Semiring R₁] [Semiring R₂] [Semiring R₃] variable [AddCommMonoid M] [AddCommMonoid M₂] [AddCommMonoid M₃] -variable [AddCommGroup N₁] [AddCommGroup N₂] [AddCommGroup N₃] +variable [AddCommGroup N₂] [AddCommGroup N₃] variable [Module R₁ M] [Module R₂ M₂] [Module R₃ M₃] -variable [Module R₁ N₁] [Module R₂ N₂] [Module R₃ N₃] +variable [Module R₂ N₂] [Module R₃ N₃] variable {σ₁₂ : R₁ →+* R₂} {σ₂₃ : R₂ →+* R₃} {σ₁₃ : R₁ →+* R₃} [RingHomCompTriple σ₁₂ σ₂₃ σ₁₃] /-- The constant 0 map is linear. -/ @@ -897,7 +894,6 @@ section SMul variable [Monoid S] [DistribMulAction S M₂] [SMulCommClass R₂ S M₂] variable [Monoid S₃] [DistribMulAction S₃ M₃] [SMulCommClass R₃ S₃ M₃] -variable [Monoid T] [DistribMulAction T M₂] [SMulCommClass R₂ T M₂] instance : DistribMulAction S (M →ₛₗ[σ₁₂] M₂) where one_smul _ := ext fun _ ↦ one_smul _ _ diff --git a/Mathlib/Algebra/Module/Pi.lean b/Mathlib/Algebra/Module/Pi.lean index 002ec529cf852..1d6a967b8db91 100644 --- a/Mathlib/Algebra/Module/Pi.lean +++ b/Mathlib/Algebra/Module/Pi.lean @@ -22,9 +22,6 @@ variable {I : Type u} -- The indexing type variable {f : I → Type v} --- The family of types already equipped with instances -variable (x y : ∀ i, f i) (i : I) - namespace Pi theorem _root_.IsSMulRegular.pi {α : Type*} [∀ i, SMul α <| f i] {k : α} diff --git a/Mathlib/Algebra/Module/Submodule/Range.lean b/Mathlib/Algebra/Module/Submodule/Range.lean index 40e7e2b452789..933cc427c94e9 100644 --- a/Mathlib/Algebra/Module/Submodule/Range.lean +++ b/Mathlib/Algebra/Module/Submodule/Range.lean @@ -28,7 +28,7 @@ linear algebra, vector space, module, range open Function variable {R : Type*} {R₂ : Type*} {R₃ : Type*} -variable {K : Type*} {K₂ : Type*} +variable {K : Type*} variable {M : Type*} {M₂ : Type*} {M₃ : Type*} variable {V : Type*} {V₂ : Type*} @@ -38,13 +38,11 @@ section AddCommMonoid variable [Semiring R] [Semiring R₂] [Semiring R₃] variable [AddCommMonoid M] [AddCommMonoid M₂] [AddCommMonoid M₃] -variable {σ₁₂ : R →+* R₂} {σ₂₃ : R₂ →+* R₃} {σ₁₃ : R →+* R₃} -variable [RingHomCompTriple σ₁₂ σ₂₃ σ₁₃] variable [Module R M] [Module R₂ M₂] [Module R₃ M₃] open Submodule -variable {σ₂₁ : R₂ →+* R} {τ₁₂ : R →+* R₂} {τ₂₃ : R₂ →+* R₃} {τ₁₃ : R →+* R₃} +variable {τ₁₂ : R →+* R₂} {τ₂₃ : R₂ →+* R₃} {τ₁₃ : R →+* R₃} variable [RingHomCompTriple τ₁₂ τ₂₃ τ₁₃] section @@ -190,11 +188,10 @@ end AddCommMonoid section Ring -variable [Ring R] [Ring R₂] [Ring R₃] -variable [AddCommGroup M] [AddCommGroup M₂] [AddCommGroup M₃] -variable [Module R M] [Module R₂ M₂] [Module R₃ M₃] -variable {τ₁₂ : R →+* R₂} {τ₂₃ : R₂ →+* R₃} {τ₁₃ : R →+* R₃} -variable [RingHomCompTriple τ₁₂ τ₂₃ τ₁₃] +variable [Ring R] [Ring R₂] +variable [AddCommGroup M] [AddCommGroup M₂] +variable [Module R M] [Module R₂ M₂] +variable {τ₁₂ : R →+* R₂} variable {F : Type*} [FunLike F M M₂] [SemilinearMapClass F τ₁₂ M M₂] variable {f : F} @@ -228,7 +225,7 @@ end Ring section Semifield -variable [Semifield K] [Semifield K₂] +variable [Semifield K] variable [AddCommMonoid V] [Module K V] variable [AddCommMonoid V₂] [Module K V₂] @@ -249,7 +246,7 @@ section AddCommMonoid variable [Semiring R] [Semiring R₂] [AddCommMonoid M] [AddCommMonoid M₂] variable [Module R M] [Module R₂ M₂] -variable (p p' : Submodule R M) (q : Submodule R₂ M₂) +variable (p : Submodule R M) variable {τ₁₂ : R →+* R₂} variable {F : Type*} [FunLike F M M₂] [SemilinearMapClass F τ₁₂ M M₂] diff --git a/Mathlib/Algebra/Order/Interval/Set/Group.lean b/Mathlib/Algebra/Order/Interval/Set/Group.lean index a95c3a44af7a1..d9684e57e4eb9 100644 --- a/Mathlib/Algebra/Order/Interval/Set/Group.lean +++ b/Mathlib/Algebra/Order/Interval/Set/Group.lean @@ -19,7 +19,7 @@ namespace Set section OrderedCommGroup -variable [OrderedCommGroup α] {a b c d : α} +variable [OrderedCommGroup α] {a c d : α} /-! `inv_mem_Ixx_iff`, `sub_mem_Ixx_iff` -/ diff --git a/Mathlib/Algebra/Order/Ring/Abs.lean b/Mathlib/Algebra/Order/Ring/Abs.lean index e61b8d4297a31..25cd393bf7434 100644 --- a/Mathlib/Algebra/Order/Ring/Abs.lean +++ b/Mathlib/Algebra/Order/Ring/Abs.lean @@ -17,7 +17,7 @@ import Mathlib.Data.Nat.Cast.Order.Ring variable {α : Type*} section LinearOrderedAddCommGroup -variable [LinearOrderedCommGroup α] {a b : α} +variable [LinearOrderedCommGroup α] @[to_additive] lemma mabs_zpow (n : ℤ) (a : α) : |a ^ n|ₘ = |a|ₘ ^ |n| := by obtain n0 | n0 := le_total 0 n @@ -34,7 +34,7 @@ lemma odd_abs [LinearOrder α] [Ring α] {a : α} : Odd (abs a) ↔ Odd a := by section LinearOrderedRing -variable [LinearOrderedRing α] {n : ℕ} {a b c : α} +variable [LinearOrderedRing α] {n : ℕ} {a b : α} @[simp] lemma abs_one : |(1 : α)| = 1 := abs_of_pos zero_lt_one @@ -141,7 +141,7 @@ end LinearOrderedRing section LinearOrderedCommRing -variable [LinearOrderedCommRing α] {a b c d : α} +variable [LinearOrderedCommRing α] theorem abs_sub_sq (a b : α) : |a - b| * |a - b| = a * a + b * b - (1 + 1) * a * b := by rw [abs_mul_abs_self] @@ -152,7 +152,7 @@ end LinearOrderedCommRing section -variable [Ring α] [LinearOrder α] {a b : α} +variable [Ring α] [LinearOrder α] @[simp] theorem abs_dvd (a b : α) : |a| ∣ b ↔ a ∣ b := by diff --git a/Mathlib/Algebra/Regular/SMul.lean b/Mathlib/Algebra/Regular/SMul.lean index 5895185e0774c..3dfb2cff80cf0 100644 --- a/Mathlib/Algebra/Regular/SMul.lean +++ b/Mathlib/Algebra/Regular/SMul.lean @@ -162,8 +162,7 @@ end MonoidSMul section MonoidWithZero -variable [MonoidWithZero R] [MonoidWithZero S] [Zero M] [MulActionWithZero R M] - [MulActionWithZero R S] [MulActionWithZero S M] [IsScalarTower R S M] +variable [MonoidWithZero R] [Zero M] [MulActionWithZero R M] /-- The element `0` is `M`-regular if and only if `M` is trivial. -/ protected theorem subsingleton (h : IsSMulRegular M (0 : R)) : Subsingleton M := diff --git a/Mathlib/Algebra/Ring/Pi.lean b/Mathlib/Algebra/Ring/Pi.lean index 394051c18c7d9..4c20d456a41b7 100644 --- a/Mathlib/Algebra/Ring/Pi.lean +++ b/Mathlib/Algebra/Ring/Pi.lean @@ -25,8 +25,7 @@ variable {I : Type u} -- The indexing type variable {f : I → Type v} --- The family of types already equipped with instances -variable (x y : ∀ i, f i) (i : I) +variable (i : I) instance distrib [∀ i, Distrib <| f i] : Distrib (∀ i : I, f i) := { add := (· + ·) diff --git a/Mathlib/Algebra/Ring/Prod.lean b/Mathlib/Algebra/Ring/Prod.lean index 2883cf4e33ea0..67ee0ebf5b125 100644 --- a/Mathlib/Algebra/Ring/Prod.lean +++ b/Mathlib/Algebra/Ring/Prod.lean @@ -23,7 +23,7 @@ trivial `simp` lemmas, and define the following operations on `RingHom`s and sim -/ -variable {α β R R' S S' T T' : Type*} +variable {R R' S S' T : Type*} namespace Prod diff --git a/Mathlib/Analysis/InnerProductSpace/Symmetric.lean b/Mathlib/Analysis/InnerProductSpace/Symmetric.lean index 9acf7b33a0635..2fcd08ccaa66d 100644 --- a/Mathlib/Analysis/InnerProductSpace/Symmetric.lean +++ b/Mathlib/Analysis/InnerProductSpace/Symmetric.lean @@ -38,11 +38,8 @@ open ComplexConjugate section Seminormed -variable {𝕜 E E' F G : Type*} [RCLike 𝕜] +variable {𝕜 E : Type*} [RCLike 𝕜] variable [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] -variable [SeminormedAddCommGroup F] [InnerProductSpace 𝕜 F] -variable [SeminormedAddCommGroup G] [InnerProductSpace 𝕜 G] -variable [SeminormedAddCommGroup E'] [InnerProductSpace ℝ E'] local notation "⟪" x ", " y "⟫" => @inner 𝕜 _ _ x y @@ -189,11 +186,8 @@ end Seminormed section Normed -variable {𝕜 E E' F G : Type*} [RCLike 𝕜] +variable {𝕜 E : Type*} [RCLike 𝕜] variable [NormedAddCommGroup E] [InnerProductSpace 𝕜 E] -variable [NormedAddCommGroup F] [InnerProductSpace 𝕜 F] -variable [NormedAddCommGroup G] [InnerProductSpace 𝕜 G] -variable [NormedAddCommGroup E'] [InnerProductSpace ℝ E'] local notation "⟪" x ", " y "⟫" => @inner 𝕜 _ _ x y diff --git a/Mathlib/Analysis/NormedSpace/OperatorNorm/Completeness.lean b/Mathlib/Analysis/NormedSpace/OperatorNorm/Completeness.lean index 783b55750d8fb..4043b6833c43d 100644 --- a/Mathlib/Analysis/NormedSpace/OperatorNorm/Completeness.lean +++ b/Mathlib/Analysis/NormedSpace/OperatorNorm/Completeness.lean @@ -20,13 +20,12 @@ open Filter hiding map_smul open scoped NNReal Topology Uniformity -- the `ₗ` subscript variables are for special cases about linear (as opposed to semilinear) maps -variable {𝕜 𝕜₂ 𝕜₃ E Eₗ F Fₗ G Gₗ 𝓕 : Type*} -variable [NormedAddCommGroup E] [NormedAddCommGroup F] [NormedAddCommGroup G] - [NormedAddCommGroup Fₗ] +variable {𝕜 𝕜₂ E F Fₗ : Type*} +variable [NormedAddCommGroup E] [NormedAddCommGroup F] [NormedAddCommGroup Fₗ] -variable [NontriviallyNormedField 𝕜] [NontriviallyNormedField 𝕜₂] [NontriviallyNormedField 𝕜₃] - [NormedSpace 𝕜 E] [NormedSpace 𝕜₂ F] [NormedSpace 𝕜₃ G] [NormedSpace 𝕜 Fₗ] (c : 𝕜) - {σ₁₂ : 𝕜 →+* 𝕜₂} {σ₂₃ : 𝕜₂ →+* 𝕜₃} (f g : E →SL[σ₁₂] F) (x y z : E) +variable [NontriviallyNormedField 𝕜] [NontriviallyNormedField 𝕜₂] + [NormedSpace 𝕜 E] [NormedSpace 𝕜₂ F] [NormedSpace 𝕜 Fₗ] + {σ₁₂ : 𝕜 →+* 𝕜₂} (f g : E →SL[σ₁₂] F) namespace ContinuousLinearMap diff --git a/Mathlib/Data/Fin/VecNotation.lean b/Mathlib/Data/Fin/VecNotation.lean index d3b561640f0ae..c931078da5ce1 100644 --- a/Mathlib/Data/Fin/VecNotation.lean +++ b/Mathlib/Data/Fin/VecNotation.lean @@ -108,7 +108,7 @@ instance _root_.PiFin.hasRepr [Repr α] : Repr (Fin n → α) where end MatrixNotation -variable {m n o : ℕ} {m' n' o' : Type*} +variable {m n o : ℕ} theorem empty_eq (v : Fin 0 → α) : v = ![] := Subsingleton.elim _ _ diff --git a/Mathlib/Data/Int/Cast/Lemmas.lean b/Mathlib/Data/Int/Cast/Lemmas.lean index da5f6ffae2ed7..822449e1e3d3b 100644 --- a/Mathlib/Data/Int/Cast/Lemmas.lean +++ b/Mathlib/Data/Int/Cast/Lemmas.lean @@ -77,7 +77,7 @@ lemma cast_ne_one : (n : α) ≠ 1 ↔ n ≠ 1 := cast_eq_one.not end AddGroupWithOne section NonAssocRing -variable [NonAssocRing α] {a b : α} {n : ℤ} +variable [NonAssocRing α] variable (α) in /-- `coe : ℤ → α` as a `RingHom`. -/ @@ -147,7 +147,7 @@ end SemiconjBy namespace Commute section NonAssocRing -variable [NonAssocRing α] {a b : α} {n : ℤ} +variable [NonAssocRing α] {a : α} {n : ℤ} @[simp] lemma intCast_left : Commute (n : α) a := Int.cast_commute _ _ @@ -159,7 +159,7 @@ variable [NonAssocRing α] {a b : α} {n : ℤ} end NonAssocRing section Ring -variable [Ring α] {a b : α} {n : ℤ} +variable [Ring α] {a b : α} @[simp] lemma intCast_mul_right (h : Commute a b) (m : ℤ) : Commute a (m * b) := SemiconjBy.intCast_mul_right h m diff --git a/Mathlib/Data/Int/Lemmas.lean b/Mathlib/Data/Int/Lemmas.lean index e989bd71419b0..3842fbb1c3222 100644 --- a/Mathlib/Data/Int/Lemmas.lean +++ b/Mathlib/Data/Int/Lemmas.lean @@ -37,8 +37,6 @@ theorem succ_natCast_pos (n : ℕ) : 0 < (n : ℤ) + 1 := /-! ### `natAbs` -/ -variable {a b : ℤ} {n : ℕ} - theorem natAbs_eq_iff_sq_eq {a b : ℤ} : a.natAbs = b.natAbs ↔ a ^ 2 = b ^ 2 := by rw [sq, sq] exact natAbs_eq_iff_mul_self_eq diff --git a/Mathlib/Data/Int/Order/Lemmas.lean b/Mathlib/Data/Int/Order/Lemmas.lean index d664a5b284aca..3e9566c7910ff 100644 --- a/Mathlib/Data/Int/Order/Lemmas.lean +++ b/Mathlib/Data/Int/Order/Lemmas.lean @@ -19,9 +19,6 @@ namespace Int /-! ### nat abs -/ - -variable {a b : ℤ} {n : ℕ} - theorem natAbs_eq_iff_mul_self_eq {a b : ℤ} : a.natAbs = b.natAbs ↔ a * a = b * b := by rw [← abs_eq_iff_mul_self_eq, abs_eq_natAbs, abs_eq_natAbs] exact Int.natCast_inj.symm diff --git a/Mathlib/Data/List/Perm.lean b/Mathlib/Data/List/Perm.lean index 8f5cb1997dd64..efeb8dd594a06 100644 --- a/Mathlib/Data/List/Perm.lean +++ b/Mathlib/Data/List/Perm.lean @@ -45,7 +45,7 @@ section Rel open Relator -variable {γ : Type*} {δ : Type*} {r : α → β → Prop} {p : γ → δ → Prop} +variable {r : α → β → Prop} local infixr:80 " ∘r " => Relation.Comp diff --git a/Mathlib/Data/List/Sort.lean b/Mathlib/Data/List/Sort.lean index 2b327d5d70ca3..513e81df6f0a3 100644 --- a/Mathlib/Data/List/Sort.lean +++ b/Mathlib/Data/List/Sort.lean @@ -49,7 +49,7 @@ namespace List section Sorted -variable {α : Type u} {β : Type v} {r : α → α → Prop} {s : β → β → Prop} {a : α} {l : List α} +variable {α : Type u} {r : α → α → Prop} {a : α} {l : List α} /-- `Sorted r l` is the same as `List.Pairwise r l`, preferred in the case that `r` is a `<` or `≤`-like relation (transitive and antisymmetric or asymmetric) -/ diff --git a/Mathlib/Data/Nat/Cast/Order/Ring.lean b/Mathlib/Data/Nat/Cast/Order/Ring.lean index b03a53d8607f4..2f6687f56c7ac 100644 --- a/Mathlib/Data/Nat/Cast/Order/Ring.lean +++ b/Mathlib/Data/Nat/Cast/Order/Ring.lean @@ -12,7 +12,7 @@ import Mathlib.Data.Nat.Cast.Order.Basic -/ -variable {α β : Type*} +variable {α : Type*} namespace Nat diff --git a/Mathlib/Dynamics/FixedPoints/Basic.lean b/Mathlib/Dynamics/FixedPoints/Basic.lean index 9f13fb19bc6cf..78081cb348764 100644 --- a/Mathlib/Dynamics/FixedPoints/Basic.lean +++ b/Mathlib/Dynamics/FixedPoints/Basic.lean @@ -27,7 +27,7 @@ open Equiv universe u v -variable {α : Type u} {β : Type v} {f fa g : α → α} {x y : α} {fb : β → β} {m n k : ℕ} {e : Perm α} +variable {α : Type u} {β : Type v} {f fa g : α → α} {x : α} {fb : β → β} {e : Perm α} namespace Function diff --git a/Mathlib/Geometry/Manifold/Algebra/LieGroup.lean b/Mathlib/Geometry/Manifold/Algebra/LieGroup.lean index 8bec2b7fa5b4f..c054be8efa60b 100644 --- a/Mathlib/Geometry/Manifold/Algebra/LieGroup.lean +++ b/Mathlib/Geometry/Manifold/Algebra/LieGroup.lean @@ -79,13 +79,10 @@ class LieGroup {𝕜 : Type*} [NontriviallyNormedField 𝕜] {H : Type*} [Topolo section PointwiseDivision variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {H : Type*} [TopologicalSpace H] {E : Type*} - [NormedAddCommGroup E] [NormedSpace 𝕜 E] {I : ModelWithCorners 𝕜 E H} {F : Type*} - [NormedAddCommGroup F] [NormedSpace 𝕜 F] {J : ModelWithCorners 𝕜 F F} {G : Type*} + [NormedAddCommGroup E] [NormedSpace 𝕜 E] {I : ModelWithCorners 𝕜 E H} {G : Type*} [TopologicalSpace G] [ChartedSpace H G] [Group G] [LieGroup I G] {E' : Type*} [NormedAddCommGroup E'] [NormedSpace 𝕜 E'] {H' : Type*} [TopologicalSpace H'] {I' : ModelWithCorners 𝕜 E' H'} {M : Type*} [TopologicalSpace M] [ChartedSpace H' M] - {E'' : Type*} [NormedAddCommGroup E''] [NormedSpace 𝕜 E''] {H'' : Type*} [TopologicalSpace H''] - {I'' : ModelWithCorners 𝕜 E'' H''} {M' : Type*} [TopologicalSpace M'] [ChartedSpace H'' M'] {n : ℕ∞} section @@ -236,7 +233,7 @@ variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {H : Type*} [TopologicalS [TopologicalSpace G] [ChartedSpace H G] [Inv G] [Zero G] [SmoothInv₀ I G] {E' : Type*} [NormedAddCommGroup E'] [NormedSpace 𝕜 E'] {H' : Type*} [TopologicalSpace H'] {I' : ModelWithCorners 𝕜 E' H'} {M : Type*} [TopologicalSpace M] [ChartedSpace H' M] - {n : ℕ∞} {f g : M → G} + {n : ℕ∞} {f : M → G} theorem smoothAt_inv₀ {x : G} (hx : x ≠ 0) : SmoothAt I I (fun y ↦ y⁻¹) x := SmoothInv₀.smoothAt_inv₀ hx diff --git a/Mathlib/Geometry/Manifold/VectorBundle/SmoothSection.lean b/Mathlib/Geometry/Manifold/VectorBundle/SmoothSection.lean index 10e4fe0a92fe3..198e5170ad73e 100644 --- a/Mathlib/Geometry/Manifold/VectorBundle/SmoothSection.lean +++ b/Mathlib/Geometry/Manifold/VectorBundle/SmoothSection.lean @@ -21,12 +21,8 @@ open Bundle Filter Function open scoped Bundle Manifold variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedAddCommGroup E] - [NormedSpace 𝕜 E] {E' : Type*} [NormedAddCommGroup E'] [NormedSpace 𝕜 E'] {H : Type*} - [TopologicalSpace H] {H' : Type*} [TopologicalSpace H'] (I : ModelWithCorners 𝕜 E H) - (I' : ModelWithCorners 𝕜 E' H') {M : Type*} [TopologicalSpace M] [ChartedSpace H M] {M' : Type*} - [TopologicalSpace M'] [ChartedSpace H' M'] {E'' : Type*} [NormedAddCommGroup E''] - [NormedSpace 𝕜 E''] {H'' : Type*} [TopologicalSpace H''] {I'' : ModelWithCorners 𝕜 E'' H''} - {M'' : Type*} [TopologicalSpace M''] [ChartedSpace H'' M''] + [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] (I : ModelWithCorners 𝕜 E H) + {M : Type*} [TopologicalSpace M] [ChartedSpace H M] variable (F : Type*) [NormedAddCommGroup F] [NormedSpace 𝕜 F] -- `F` model fiber @@ -51,7 +47,7 @@ abbrev SmoothSection := namespace ContMDiffSection -variable {I} {I'} {n} {F} {V} +variable {I} {n} {F} {V} instance : DFunLike Cₛ^n⟮I; F, V⟯ M V where coe := ContMDiffSection.toFun diff --git a/Mathlib/GroupTheory/Congruence/Defs.lean b/Mathlib/GroupTheory/Congruence/Defs.lean index c8c38842eee77..ced9c0f1660d1 100644 --- a/Mathlib/GroupTheory/Congruence/Defs.lean +++ b/Mathlib/GroupTheory/Congruence/Defs.lean @@ -584,7 +584,7 @@ end section MulOneClass -variable [MulOneClass M] [MulOneClass N] [MulOneClass P] (c : Con M) +variable [MulOneClass M] (c : Con M) /-- The quotient of a monoid by a congruence relation is a monoid. -/ @[to_additive "The quotient of an `AddMonoid` by an additive congruence relation is @@ -709,7 +709,7 @@ end Monoids section Groups -variable [Group M] [Group N] [Group P] (c : Con M) +variable [Group M] (c : Con M) /-- Multiplicative congruence relations preserve inversion. -/ @[to_additive "Additive congruence relations preserve negation."] diff --git a/Mathlib/GroupTheory/CosetCover.lean b/Mathlib/GroupTheory/CosetCover.lean index 03164c5973c5e..a8ba591868ffa 100644 --- a/Mathlib/GroupTheory/CosetCover.lean +++ b/Mathlib/GroupTheory/CosetCover.lean @@ -366,7 +366,7 @@ end Submodule section Subspace -variable {k E ι : Type*} [DivisionRing k] [Infinite k] [AddCommGroup E] [Module k E] +variable {k E : Type*} [DivisionRing k] [Infinite k] [AddCommGroup E] [Module k E] {s : Finset (Subspace k E)} /- A vector space over an infinite field cannot be a finite union of proper subspaces. -/ diff --git a/Mathlib/GroupTheory/DoubleCoset.lean b/Mathlib/GroupTheory/DoubleCoset.lean index 771be71a00bb9..a7d7e29638510 100644 --- a/Mathlib/GroupTheory/DoubleCoset.lean +++ b/Mathlib/GroupTheory/DoubleCoset.lean @@ -22,7 +22,7 @@ this is the usual left or right quotient of a group by a subgroup. -- Porting note: removed import -- import Mathlib.Tactic.Group -variable {G : Type*} [Group G] {α : Type*} [Mul α] (J : Subgroup G) (g : G) +variable {G : Type*} [Group G] {α : Type*} [Mul α] open MulOpposite open scoped Pointwise diff --git a/Mathlib/GroupTheory/GroupAction/Hom.lean b/Mathlib/GroupTheory/GroupAction/Hom.lean index e2fc5be8f79d8..fdebbaf6ec0f9 100644 --- a/Mathlib/GroupTheory/GroupAction/Hom.lean +++ b/Mathlib/GroupTheory/GroupAction/Hom.lean @@ -524,13 +524,10 @@ def inverse (f : A →+[M] B₁) (g : B₁ → A) (h₁ : Function.LeftInverse g section Semiring variable (R : Type*) [Semiring R] [MulSemiringAction M R] -variable (R' : Type*) [Ring R'] [MulSemiringAction M R'] variable (S : Type*) [Semiring S] [MulSemiringAction N S] -variable (S' : Type*) [Ring S'] [MulSemiringAction N S'] variable (T : Type*) [Semiring T] [MulSemiringAction P T] -variable {R S M' N'} -variable [AddMonoid M'] [DistribMulAction R M'] +variable {R S N'} variable [AddMonoid N'] [DistribMulAction S N'] variable {σ : R →* S} diff --git a/Mathlib/GroupTheory/NoncommPiCoprod.lean b/Mathlib/GroupTheory/NoncommPiCoprod.lean index 70be3543599df..d675c2b73a7f5 100644 --- a/Mathlib/GroupTheory/NoncommPiCoprod.lean +++ b/Mathlib/GroupTheory/NoncommPiCoprod.lean @@ -90,9 +90,6 @@ variable (ϕ : ∀ i : ι, N i →* M) -- We assume that the elements of different morphism commute variable (hcomm : Pairwise fun i j => ∀ x y, Commute (ϕ i x) (ϕ j y)) --- We use `f` and `g` to denote elements of `Π (i : ι), N i` -variable (f g : ∀ i : ι, N i) - namespace MonoidHom /-- The canonical homomorphism from a family of monoids. -/ @@ -168,9 +165,6 @@ variable {ι : Type*} variable {H : ι → Type*} [∀ i, Group (H i)] variable (ϕ : ∀ i : ι, H i →* G) --- We use `f` and `g` to denote elements of `Π (i : ι), H i` -variable (f g : ∀ i : ι, H i) - namespace MonoidHom -- The subgroup version of `MonoidHom.noncommPiCoprod_mrange` @[to_additive] @@ -251,9 +245,6 @@ namespace Subgroup variable {G : Type*} [Group G] variable {ι : Type*} {H : ι → Subgroup G} --- Elements of `Π (i : ι), H i` are called `f` and `g` here -variable (f g : ∀ i : ι, H i) - section CommutingSubgroups -- We assume that the elements of different subgroups commute diff --git a/Mathlib/GroupTheory/Sylow.lean b/Mathlib/GroupTheory/Sylow.lean index 7792a01f89afb..6609de9811ac4 100644 --- a/Mathlib/GroupTheory/Sylow.lean +++ b/Mathlib/GroupTheory/Sylow.lean @@ -414,9 +414,9 @@ end InfiniteSylow open Equiv Equiv.Perm Finset Function List QuotientGroup -universe u v w +universe u -variable {G : Type u} {α : Type v} {β : Type w} [Group G] +variable {G : Type u} [Group G] theorem QuotientGroup.card_preimage_mk (s : Subgroup G) (t : Set (G ⧸ s)) : Nat.card (QuotientGroup.mk ⁻¹' t) = Nat.card s * Nat.card t := by diff --git a/Mathlib/LinearAlgebra/Basis/Basic.lean b/Mathlib/LinearAlgebra/Basis/Basic.lean index 46ed8282bcd9a..3e294300107a0 100644 --- a/Mathlib/LinearAlgebra/Basis/Basic.lean +++ b/Mathlib/LinearAlgebra/Basis/Basic.lean @@ -24,18 +24,15 @@ universe u open Function Set Submodule Finsupp -variable {ι : Type*} {ι' : Type*} {R : Type*} {R₂ : Type*} {K : Type*} -variable {M : Type*} {M' M'' : Type*} {V : Type u} {V' : Type*} +variable {ι : Type*} {ι' : Type*} {R : Type*} {R₂ : Type*} {M : Type*} {M' : Type*} section Module -variable [Semiring R] -variable [AddCommMonoid M] [Module R M] [AddCommMonoid M'] [Module R M'] - +variable [Semiring R] [AddCommMonoid M] [Module R M] [AddCommMonoid M'] [Module R M'] namespace Basis -variable (b b₁ : Basis ι R M) (i : ι) (c : R) (x : M) +variable (b : Basis ι R M) section Coord @@ -187,9 +184,9 @@ section Module open LinearMap variable {v : ι → M} -variable [Ring R] [CommRing R₂] [AddCommGroup M] [AddCommGroup M'] [AddCommGroup M''] -variable [Module R M] [Module R₂ M] [Module R M'] [Module R M''] -variable {c d : R} {x y : M} +variable [Ring R] [CommRing R₂] [AddCommGroup M] +variable [Module R M] [Module R₂ M] +variable {x y : M} variable (b : Basis ι R M) namespace Basis diff --git a/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean b/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean index 14b5f34c2a2b3..172f5f1c0d87a 100644 --- a/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean +++ b/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean @@ -162,8 +162,7 @@ end FiniteDimensional namespace Module variable (K V) -variable [DivisionRing K] [AddCommGroup V] [Module K V] {V₂ : Type v'} [AddCommGroup V₂] - [Module K V₂] +variable [DivisionRing K] [AddCommGroup V] [Module K V] /-- In a finite-dimensional space, its dimension (seen as a cardinal) coincides with its `finrank`. This is a copy of `finrank_eq_rank _ _` which creates easier typeclass searches. -/ @@ -426,9 +425,6 @@ protected theorem finiteDimensional (f : V ≃ₗ[K] V₂) [FiniteDimensional K FiniteDimensional K V₂ := Module.Finite.equiv f -variable {R M M₂ : Type*} [Ring R] [AddCommGroup M] [AddCommGroup M₂] -variable [Module R M] [Module R M₂] - end LinearEquiv section @@ -442,8 +438,7 @@ instance finiteDimensional_finsupp {ι : Type*} [Finite ι] [FiniteDimensional K end namespace Submodule -variable [DivisionRing K] [AddCommGroup V] [Module K V] {V₂ : Type v'} [AddCommGroup V₂] - [Module K V₂] +variable [DivisionRing K] [AddCommGroup V] [Module K V] /-- If a submodule is contained in a finite-dimensional submodule with the same or smaller dimension, they are equal. -/ diff --git a/Mathlib/MeasureTheory/Integral/Bochner.lean b/Mathlib/MeasureTheory/Integral/Bochner.lean index 411b231ffbadc..b4eab8d959db6 100644 --- a/Mathlib/MeasureTheory/Integral/Bochner.lean +++ b/Mathlib/MeasureTheory/Integral/Bochner.lean @@ -279,9 +279,8 @@ and prove basic property of this integral. open Finset -variable [NormedAddCommGroup E] [NormedAddCommGroup F] [NormedSpace ℝ F] {p : ℝ≥0∞} {G F' : Type*} - [NormedAddCommGroup G] [NormedAddCommGroup F'] [NormedSpace ℝ F'] {m : MeasurableSpace α} - {μ : Measure α} +variable [NormedAddCommGroup E] [NormedAddCommGroup F] [NormedSpace ℝ F] + {m : MeasurableSpace α} {μ : Measure α} /-- Bochner integral of simple functions whose codomain is a real `NormedSpace`. This is equal to `∑ x ∈ f.range, (μ (f ⁻¹' {x})).toReal • x` (see `integral_eq`). -/ @@ -423,7 +422,7 @@ namespace L1 open AEEqFun Lp.simpleFunc Lp -variable [NormedAddCommGroup E] [NormedAddCommGroup F] {m : MeasurableSpace α} {μ : Measure α} +variable [NormedAddCommGroup E] {m : MeasurableSpace α} {μ : Measure α} namespace SimpleFunc @@ -464,8 +463,7 @@ Define the Bochner integral on `α →₁ₛ[μ] E` by extension from the simple and prove basic properties of this integral. -/ -variable [NormedField 𝕜] [NormedSpace 𝕜 E] [NormedSpace ℝ E] [SMulCommClass ℝ 𝕜 E] {F' : Type*} - [NormedAddCommGroup F'] [NormedSpace ℝ F'] +variable [NormedField 𝕜] [NormedSpace 𝕜 E] [NormedSpace ℝ E] [SMulCommClass ℝ 𝕜 E] attribute [local instance] simpleFunc.normedSpace @@ -495,7 +493,6 @@ theorem norm_integral_le_norm (f : α →₁ₛ[μ] E) : ‖integral f‖ ≤ rw [integral, norm_eq_integral] exact (toSimpleFunc f).norm_integral_le_integral_norm (SimpleFunc.integrable f) -variable {E' : Type*} [NormedAddCommGroup E'] [NormedSpace ℝ E'] [NormedSpace 𝕜 E'] variable (α E μ 𝕜) /-- The Bochner integral over simple functions in L1 space as a continuous linear map. -/ @@ -575,7 +572,7 @@ open SimpleFunc local notation "Integral" => @integralCLM α E _ _ _ _ _ μ _ variable [NormedSpace ℝ E] [NontriviallyNormedField 𝕜] [NormedSpace 𝕜 E] [SMulCommClass ℝ 𝕜 E] - [NormedSpace ℝ F] [CompleteSpace E] + [CompleteSpace E] section IntegrationInL1 @@ -700,7 +697,7 @@ functions, and 0 otherwise; prove its basic properties. -/ variable [NormedAddCommGroup E] [hE : CompleteSpace E] [NontriviallyNormedField 𝕜] - [NormedSpace 𝕜 E] [NormedAddCommGroup F] [NormedSpace ℝ F] [CompleteSpace F] + [NormedAddCommGroup F] [NormedSpace ℝ F] [CompleteSpace F] {G : Type*} [NormedAddCommGroup G] [NormedSpace ℝ G] open Classical in @@ -730,8 +727,8 @@ section Properties open ContinuousLinearMap MeasureTheory.SimpleFunc -variable [NormedSpace ℝ E] [SMulCommClass ℝ 𝕜 E] -variable {f g : α → E} {m : MeasurableSpace α} {μ : Measure α} +variable [NormedSpace ℝ E] +variable {f : α → E} {m : MeasurableSpace α} {μ : Measure α} theorem integral_eq (f : α → E) (hf : Integrable f μ) : ∫ a, f a ∂μ = L1.integral (hf.toL1 f) := by simp [integral, hE, hf] @@ -1788,7 +1785,7 @@ end Properties section IntegralTrim -variable {H β γ : Type*} [NormedAddCommGroup H] {m m0 : MeasurableSpace β} {μ : Measure β} +variable {β γ : Type*} {m m0 : MeasurableSpace β} {μ : Measure β} /-- Simple function seen as simple function of a larger `MeasurableSpace`. -/ def SimpleFunc.toLargerSpace (hm : m ≤ m0) (f : @SimpleFunc β m γ) : SimpleFunc β γ := diff --git a/Mathlib/Order/Bounds/Basic.lean b/Mathlib/Order/Bounds/Basic.lean index eaee5e41d1330..fb6e31a954dff 100644 --- a/Mathlib/Order/Bounds/Basic.lean +++ b/Mathlib/Order/Bounds/Basic.lean @@ -26,7 +26,7 @@ variable {α : Type u} {β : Type v} {γ : Type w} {ι : Sort x} section -variable [Preorder α] [Preorder β] {s t : Set α} {a b : α} +variable [Preorder α] {s t : Set α} {a b : α} theorem mem_upperBounds : a ∈ upperBounds s ↔ ∀ x ∈ s, x ≤ a := Iff.rfl diff --git a/Mathlib/Order/CompleteLattice.lean b/Mathlib/Order/CompleteLattice.lean index dc192f4a528ec..2d47e7b97ad07 100644 --- a/Mathlib/Order/CompleteLattice.lean +++ b/Mathlib/Order/CompleteLattice.lean @@ -45,7 +45,7 @@ In lemma names, open Function OrderDual Set -variable {α β β₂ γ : Type*} {ι ι' : Sort*} {κ : ι → Sort*} {κ' : ι' → Sort*} +variable {α β γ : Type*} {ι ι' : Sort*} {κ : ι → Sort*} {κ' : ι' → Sort*} @[simp] lemma iSup_ulift {ι : Type*} [SupSet α] (f : ULift ι → α) : ⨆ i : ULift ι, f i = ⨆ i, f (.up i) := by simp [iSup]; congr with x; simp diff --git a/Mathlib/Order/GaloisConnection.lean b/Mathlib/Order/GaloisConnection.lean index d4ea55c673a62..7420bcaf290d1 100644 --- a/Mathlib/Order/GaloisConnection.lean +++ b/Mathlib/Order/GaloisConnection.lean @@ -50,8 +50,8 @@ open Function OrderDual Set universe u v w x -variable {α : Type u} {β : Type v} {γ : Type w} {ι : Sort x} {κ : ι → Sort*} {a a₁ a₂ : α} - {b b₁ b₂ : β} +variable {α : Type u} {β : Type v} {γ : Type w} {ι : Sort x} {κ : ι → Sort*} {a₁ a₂ : α} + {b₁ b₂ : β} /-- A Galois connection is a pair of functions `l` and `u` satisfying `l a ≤ b ↔ a ≤ u b`. They are special cases of adjoint functors in category theory, @@ -337,7 +337,7 @@ theorem gc_Ici_sInf [CompleteSemilatticeInf α] : GaloisConnection (toDual ∘ Ici : α → (Set α)ᵒᵈ) (sInf ∘ ofDual : (Set α)ᵒᵈ → α) := fun _ _ ↦ le_sInf_iff.symm -variable [CompleteLattice α] [CompleteLattice β] [CompleteLattice γ] {f : α → β → γ} {s : Set α} +variable [CompleteLattice α] [CompleteLattice β] [CompleteLattice γ] {s : Set α} {t : Set β} {l u : α → β → γ} {l₁ u₁ : β → γ → α} {l₂ u₂ : α → γ → β} theorem sSup_image2_eq_sSup_sSup (h₁ : ∀ b, GaloisConnection (swap l b) (u₁ b)) diff --git a/Mathlib/Order/PiLex.lean b/Mathlib/Order/PiLex.lean index bcad74a66ca1d..4fd435168435e 100644 --- a/Mathlib/Order/PiLex.lean +++ b/Mathlib/Order/PiLex.lean @@ -101,7 +101,7 @@ noncomputable instance [LinearOrder ι] [IsWellOrder ι (· < ·)] [∀ a, Linea section PartialOrder -variable [LinearOrder ι] [IsWellOrder ι (· < ·)] [∀ i, PartialOrder (β i)] {x y : ∀ i, β i} {i : ι} +variable [LinearOrder ι] [IsWellOrder ι (· < ·)] [∀ i, PartialOrder (β i)] {x : ∀ i, β i} {i : ι} {a : β i} open Function diff --git a/Mathlib/RingTheory/Localization/Submodule.lean b/Mathlib/RingTheory/Localization/Submodule.lean index 64d723f0fd50d..1e3d73e6d705b 100644 --- a/Mathlib/RingTheory/Localization/Submodule.lean +++ b/Mathlib/RingTheory/Localization/Submodule.lean @@ -21,7 +21,7 @@ commutative ring, field of fractions variable {R : Type*} [CommRing R] (M : Submonoid R) (S : Type*) [CommRing S] -variable [Algebra R S] {P : Type*} [CommRing P] +variable [Algebra R S] namespace IsLocalization @@ -72,9 +72,6 @@ theorem coeSubmodule_span_singleton (x : R) : coeSubmodule S (Ideal.span {x}) = Submodule.span R {(algebraMap R S) x} := by rw [coeSubmodule_span, Set.image_singleton] -variable {g : R →+* P} -variable {T : Submonoid P} (hy : M ≤ T.comap g) {Q : Type*} [CommRing Q] -variable [Algebra P Q] [IsLocalization T Q] variable [IsLocalization M S] include M in @@ -159,11 +156,11 @@ namespace IsFractionRing open IsLocalization -variable {A K : Type*} [CommRing A] +variable {K : Type*} section CommRing -variable [CommRing K] [Algebra R K] [IsFractionRing R K] [Algebra A K] [IsFractionRing A K] +variable [CommRing K] [Algebra R K] [IsFractionRing R K] @[simp, mono] theorem coeSubmodule_le_coeSubmodule {I J : Ideal R} : diff --git a/Mathlib/Topology/Algebra/Module/Basic.lean b/Mathlib/Topology/Algebra/Module/Basic.lean index df022af891410..637ef9b981a96 100644 --- a/Mathlib/Topology/Algebra/Module/Basic.lean +++ b/Mathlib/Topology/Algebra/Module/Basic.lean @@ -130,9 +130,8 @@ end Submodule section closure -variable {R R' : Type u} {M M' : Type v} [Semiring R] [Ring R'] - [TopologicalSpace M] [AddCommMonoid M] [TopologicalSpace M'] [AddCommGroup M'] [Module R M] - [ContinuousConstSMul R M] [Module R' M'] [ContinuousConstSMul R' M'] +variable {R : Type u} {M : Type v} [Semiring R] [TopologicalSpace M] [AddCommMonoid M] [Module R M] + [ContinuousConstSMul R M] theorem Submodule.mapsTo_smul_closure (s : Submodule R M) (c : R) : Set.MapsTo (c • ·) (closure s : Set M) (closure s) := @@ -1418,7 +1417,7 @@ variable {R R₂ R₃ S S₃ : Type*} [Semiring R] [Semiring R₂] [Semiring R [SMulCommClass R₃ S₃ M₃] [ContinuousConstSMul S₃ M₃] [Module S N₂] [ContinuousConstSMul S N₂] [SMulCommClass R S N₂] [Module S N₃] [SMulCommClass R S N₃] [ContinuousConstSMul S N₃] {σ₁₂ : R →+* R₂} {σ₂₃ : R₂ →+* R₃} {σ₁₃ : R →+* R₃} [RingHomCompTriple σ₁₂ σ₂₃ σ₁₃] (c : S) - (h : M₂ →SL[σ₂₃] M₃) (f g : M →SL[σ₁₂] M₂) (x y z : M) + (h : M₂ →SL[σ₂₃] M₃) (f : M →SL[σ₁₂] M₂) /-- `ContinuousLinearMap.prod` as an `Equiv`. -/ @[simps apply] @@ -1505,7 +1504,7 @@ section CommRing variable {R : Type*} [CommRing R] {M : Type*} [TopologicalSpace M] [AddCommGroup M] {M₂ : Type*} [TopologicalSpace M₂] [AddCommGroup M₂] {M₃ : Type*} [TopologicalSpace M₃] [AddCommGroup M₃] - [Module R M] [Module R M₂] [Module R M₃] [ContinuousConstSMul R M₃] + [Module R M] [Module R M₂] [Module R M₃] variable [TopologicalAddGroup M₂] [ContinuousConstSMul R M₂] @@ -1596,10 +1595,10 @@ variable {R₁ : Type*} {R₂ : Type*} {R₃ : Type*} [Semiring R₁] [Semiring {σ₂₃ : R₂ →+* R₃} {σ₃₂ : R₃ →+* R₂} [RingHomInvPair σ₂₃ σ₃₂] [RingHomInvPair σ₃₂ σ₂₃] {σ₁₃ : R₁ →+* R₃} {σ₃₁ : R₃ →+* R₁} [RingHomInvPair σ₁₃ σ₃₁] [RingHomInvPair σ₃₁ σ₁₃] [RingHomCompTriple σ₁₂ σ₂₃ σ₁₃] [RingHomCompTriple σ₃₂ σ₂₁ σ₃₁] {M₁ : Type*} - [TopologicalSpace M₁] [AddCommMonoid M₁] {M'₁ : Type*} [TopologicalSpace M'₁] [AddCommMonoid M'₁] + [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type*} [TopologicalSpace M₂] [AddCommMonoid M₂] {M₃ : Type*} [TopologicalSpace M₃] [AddCommMonoid M₃] {M₄ : Type*} [TopologicalSpace M₄] [AddCommMonoid M₄] [Module R₁ M₁] - [Module R₁ M'₁] [Module R₂ M₂] [Module R₃ M₃] + [Module R₂ M₂] [Module R₃ M₃] /-- A continuous linear equivalence induces a continuous linear map. -/ @[coe] @@ -2302,7 +2301,6 @@ end ContinuousLinearMap namespace Submodule variable {R : Type*} [Ring R] {M : Type*} [TopologicalSpace M] [AddCommGroup M] [Module R M] - {M₂ : Type*} [TopologicalSpace M₂] [AddCommGroup M₂] [Module R M₂] open ContinuousLinearMap From c744def559485980adb40b991247176a127345b7 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 17 Oct 2024 16:19:09 +0000 Subject: [PATCH 105/505] feat(Algebra/Module/Submodule): lemmas about `domRestrict` (#17806) From PFR --- Mathlib/Algebra/Module/Submodule/Ker.lean | 10 +++++++--- Mathlib/Algebra/Module/Submodule/Map.lean | 5 ++++- Mathlib/Algebra/Module/Submodule/Range.lean | 3 +++ 3 files changed, 14 insertions(+), 4 deletions(-) diff --git a/Mathlib/Algebra/Module/Submodule/Ker.lean b/Mathlib/Algebra/Module/Submodule/Ker.lean index 23ce2537e64cc..b291dd847675d 100644 --- a/Mathlib/Algebra/Module/Submodule/Ker.lean +++ b/Mathlib/Algebra/Module/Submodule/Ker.lean @@ -109,10 +109,13 @@ theorem le_ker_iff_map [RingHomSurjective τ₁₂] {f : F} {p : Submodule R M} theorem ker_codRestrict {τ₂₁ : R₂ →+* R} (p : Submodule R M) (f : M₂ →ₛₗ[τ₂₁] M) (hf) : ker (codRestrict p f hf) = ker f := by rw [ker, comap_codRestrict, Submodule.map_bot]; rfl +lemma ker_domRestrict [AddCommMonoid M₁] [Module R M₁] (p : Submodule R M) (f : M →ₗ[R] M₁) : + ker (domRestrict f p) = (ker f).comap p.subtype := ker_comp .. + theorem ker_restrict [AddCommMonoid M₁] [Module R M₁] {p : Submodule R M} {q : Submodule R M₁} {f : M →ₗ[R] M₁} (hf : ∀ x : M, x ∈ p → f x ∈ q) : - ker (f.restrict hf) = LinearMap.ker (f.domRestrict p) := by - rw [restrict_eq_codRestrict_domRestrict, ker_codRestrict] + ker (f.restrict hf) = (ker f).comap p.subtype := by + rw [restrict_eq_codRestrict_domRestrict, ker_codRestrict, ker_domRestrict] @[simp] theorem ker_zero : ker (0 : M →ₛₗ[τ₁₂] M₂) = ⊤ := @@ -198,7 +201,8 @@ theorem ker_eq_bot {f : M →ₛₗ[τ₁₂] M₂} : ker f = ⊥ ↔ Injective @[simp] theorem injective_restrict_iff_disjoint {p : Submodule R M} {f : M →ₗ[R] M} (hf : ∀ x ∈ p, f x ∈ p) : Injective (f.restrict hf) ↔ Disjoint p (ker f) := by - rw [← ker_eq_bot, ker_restrict hf, ker_eq_bot, injective_domRestrict_iff, disjoint_iff] + rw [← ker_eq_bot, ker_restrict hf, ← ker_domRestrict, ker_eq_bot, injective_domRestrict_iff, + disjoint_iff] end Ring diff --git a/Mathlib/Algebra/Module/Submodule/Map.lean b/Mathlib/Algebra/Module/Submodule/Map.lean index 8e6b50ba4e465..84ea00362c6d2 100644 --- a/Mathlib/Algebra/Module/Submodule/Map.lean +++ b/Mathlib/Algebra/Module/Submodule/Map.lean @@ -290,9 +290,12 @@ theorem map_iInf_comap_of_surjective {ι : Sort*} (S : ι → Submodule R₂ M (⨅ i, (S i).comap f).map f = iInf S := (giMapComap hf).l_iInf_u _ -theorem comap_le_comap_iff_of_surjective (p q : Submodule R₂ M₂) : p.comap f ≤ q.comap f ↔ p ≤ q := +theorem comap_le_comap_iff_of_surjective {p q : Submodule R₂ M₂} : p.comap f ≤ q.comap f ↔ p ≤ q := (giMapComap hf).u_le_u_iff +lemma comap_lt_comap_iff_of_surjective {p q : Submodule R₂ M₂} : p.comap f < q.comap f ↔ p < q := by + apply lt_iff_lt_of_le_iff_le' <;> exact comap_le_comap_iff_of_surjective hf + theorem comap_strictMono_of_surjective : StrictMono (comap f) := (giMapComap hf).strictMono_u diff --git a/Mathlib/Algebra/Module/Submodule/Range.lean b/Mathlib/Algebra/Module/Submodule/Range.lean index 933cc427c94e9..1053c1cf8c54c 100644 --- a/Mathlib/Algebra/Module/Submodule/Range.lean +++ b/Mathlib/Algebra/Module/Submodule/Range.lean @@ -100,6 +100,9 @@ theorem range_neg {R : Type*} {R₂ : Type*} {M : Type*} {M₂ : Type*} [Semirin change range ((-LinearMap.id : M₂ →ₗ[R₂] M₂).comp f) = _ rw [range_comp, Submodule.map_neg, Submodule.map_id] +@[simp] lemma range_domRestrict [Module R M₂] (K : Submodule R M) (f : M →ₗ[R] M₂) : + range (domRestrict f K) = K.map f := by ext; simp + lemma range_domRestrict_le_range [RingHomSurjective τ₁₂] (f : M →ₛₗ[τ₁₂] M₂) (S : Submodule R M) : LinearMap.range (f.domRestrict S) ≤ LinearMap.range f := by rintro x ⟨⟨y, hy⟩, rfl⟩ From 9d883e2896b1b16d8e646c336c647da4fc1cacd5 Mon Sep 17 00:00:00 2001 From: FMLJohn <417963616@qq.com> Date: Thu, 17 Oct 2024 17:56:40 +0000 Subject: [PATCH 106/505] feat(RingTheory/PowerSeries/WellKnown): changed `PowerSeries.invOneSubPow` (#17360) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The definition of `PowerSeries.invOneSubPow` has been changed into the following: Given a commutative ring `S` and a number `d : ℕ`, `PowerSeries.invOneSubPow S d` is the multiplicative inverse of `(1 - X) ^ d` in `S⟦X⟧ˣ`. When `d` is `0`, `PowerSeries.invOneSubPow S d` will just be `1`. When `d` is positive, `PowerSeries.invOneSubPow S d` will be `∑ n, Nat.choose (d - 1 + n) (d - 1)`. Co-authored-by: FMLJohn <128468682+FMLJohn@users.noreply.github.com> --- Mathlib/RingTheory/PowerSeries/WellKnown.lean | 82 +++++++++++++------ 1 file changed, 56 insertions(+), 26 deletions(-) diff --git a/Mathlib/RingTheory/PowerSeries/WellKnown.lean b/Mathlib/RingTheory/PowerSeries/WellKnown.lean index 8b3ce3c34a0ce..292659bb8862f 100644 --- a/Mathlib/RingTheory/PowerSeries/WellKnown.lean +++ b/Mathlib/RingTheory/PowerSeries/WellKnown.lean @@ -18,8 +18,9 @@ In this file we define the following power series: It is given by `∑ n, x ^ n /ₚ u ^ (n + 1)`. * `PowerSeries.invOneSubPow`: given a commutative ring `S` and a number `d : ℕ`, - `PowerSeries.invOneSubPow d : S⟦X⟧ˣ` is the power series `∑ n, Nat.choose (d + n) d` - whose multiplicative inverse is `(1 - X) ^ (d + 1)`. + `PowerSeries.invOneSubPow S d` is the multiplicative inverse of `(1 - X) ^ d` in `S⟦X⟧ˣ`. + When `d` is `0`, `PowerSeries.invOneSubPow S d` will just be `1`. When `d` is positive, + `PowerSeries.invOneSubPow S d` will be `∑ n, Nat.choose (d - 1 + n) (d - 1)`. * `PowerSeries.sin`, `PowerSeries.cos`, `PowerSeries.exp` : power series for sin, cosine, and exponential functions. @@ -64,7 +65,7 @@ end Ring section invOneSubPow -variable {S : Type*} [CommRing S] (d : ℕ) +variable (S : Type*) [CommRing S] (d : ℕ) /-- (1 + X + X^2 + ...) * (1 - X) = 1. @@ -97,39 +98,68 @@ theorem mk_one_pow_eq_mk_choose_add : add_right_comm] /-- -The power series `mk fun n => Nat.choose (d + n) d`, whose multiplicative inverse is -`(1 - X) ^ (d + 1)`. +Given a natural number `d : ℕ` and a commutative ring `S`, `PowerSeries.invOneSubPow S d` is the +multiplicative inverse of `(1 - X) ^ d` in `S⟦X⟧ˣ`. When `d` is `0`, `PowerSeries.invOneSubPow S d` +will just be `1`. When `d` is positive, `PowerSeries.invOneSubPow S d` will be the power series +`mk fun n => Nat.choose (d - 1 + n) (d - 1)`. -/ -noncomputable def invOneSubPow : S⟦X⟧ˣ where - val := mk fun n => Nat.choose (d + n) d - inv := (1 - X) ^ (d + 1) - val_inv := by - rw [← mk_one_pow_eq_mk_choose_add, ← mul_pow, mk_one_mul_one_sub_eq_one, one_pow] - inv_val := by - rw [← mk_one_pow_eq_mk_choose_add, ← mul_pow, mul_comm, mk_one_mul_one_sub_eq_one, one_pow] - -theorem invOneSubPow_val_eq_mk_choose_add : - (invOneSubPow d).val = (mk fun n => Nat.choose (d + n) d : S⟦X⟧) := rfl - -theorem invOneSubPow_val_zero_eq_invUnitSub_one : - (invOneSubPow 0).val = invUnitsSub (1 : Sˣ) := by +noncomputable def invOneSubPow : ℕ → S⟦X⟧ˣ + | 0 => 1 + | d + 1 => { + val := mk fun n => Nat.choose (d + n) d + inv := (1 - X) ^ (d + 1) + val_inv := by + rw [← mk_one_pow_eq_mk_choose_add, ← mul_pow, mk_one_mul_one_sub_eq_one, one_pow] + inv_val := by + rw [← mk_one_pow_eq_mk_choose_add, ← mul_pow, mul_comm, mk_one_mul_one_sub_eq_one, one_pow] + } + +theorem invOneSubPow_zero : invOneSubPow S 0 = 1 := by + delta invOneSubPow + simp only [Units.val_one] + +theorem invOneSubPow_val_eq_mk_sub_one_add_choose_of_pos (h : 0 < d) : + (invOneSubPow S d).val = (mk fun n => Nat.choose (d - 1 + n) (d - 1) : S⟦X⟧) := by + rw [← Nat.sub_one_add_one_eq_of_pos h, invOneSubPow, add_tsub_cancel_right] + +theorem invOneSubPow_val_succ_eq_mk_add_choose : + (invOneSubPow S (d + 1)).val = (mk fun n => Nat.choose (d + n) d : S⟦X⟧) := rfl + +theorem invOneSubPow_val_one_eq_invUnitSub_one : + (invOneSubPow S 1).val = invUnitsSub (1 : Sˣ) := by simp [invOneSubPow, invUnitsSub] /-- The theorem `PowerSeries.mk_one_mul_one_sub_eq_one` implies that `1 - X` is a unit in `S⟦X⟧` whose inverse is the power series `1 + X + X^2 + ...`. This theorem states that for any `d : ℕ`, -`PowerSeries.invOneSubPow d` is equal to `(1 - X)⁻¹ ^ (d + 1)`. +`PowerSeries.invOneSubPow S d` is equal to `(1 - X)⁻¹ ^ d`. -/ theorem invOneSubPow_eq_inv_one_sub_pow : - invOneSubPow d = (Units.mkOfMulEqOne (1 - X) (mk 1 : S⟦X⟧) - <| Eq.trans (mul_comm _ _) mk_one_mul_one_sub_eq_one)⁻¹ ^ (d + 1) := by - rw [inv_pow] - exact (DivisionMonoid.inv_eq_of_mul _ (invOneSubPow d) <| by - rw [← Units.val_eq_one, Units.val_mul, Units.val_pow_eq_pow_val] - exact (invOneSubPow d).inv_val).symm + invOneSubPow S d = + (Units.mkOfMulEqOne (1 - X) (mk 1 : S⟦X⟧) <| + Eq.trans (mul_comm _ _) (mk_one_mul_one_sub_eq_one S))⁻¹ ^ d := by + induction d with + | zero => exact Eq.symm <| pow_zero _ + | succ d _ => + rw [inv_pow] + exact (DivisionMonoid.inv_eq_of_mul _ (invOneSubPow S (d + 1)) <| by + rw [← Units.val_eq_one, Units.val_mul, Units.val_pow_eq_pow_val] + exact (invOneSubPow S (d + 1)).inv_val).symm theorem invOneSubPow_inv_eq_one_sub_pow : - (invOneSubPow d).inv = (1 - X : S⟦X⟧) ^ (d + 1) := rfl + (invOneSubPow S d).inv = (1 - X : S⟦X⟧) ^ d := by + induction d with + | zero => exact Eq.symm <| pow_zero _ + | succ d => rfl + +theorem invOneSubPow_inv_eq_one_of_eq_zero (h : d = 0) : + (invOneSubPow S d).inv = 1 := by + delta invOneSubPow + simp only [h, Units.inv_eq_val_inv, inv_one, Units.val_one] + +theorem mk_add_choose_mul_one_sub_pow_eq_one : + (mk fun n ↦ Nat.choose (d + n) d : S⟦X⟧) * ((1 - X) ^ (d + 1)) = 1 := + (invOneSubPow S (d + 1)).val_inv end invOneSubPow From e759b3b1e909c1ebe13a364745f063152a61a3aa Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Thu, 17 Oct 2024 19:04:09 +0000 Subject: [PATCH 107/505] feat: UniqueFactorizationMonoid.factors_rel_of_associated (#17260) --- Mathlib/RingTheory/Radical.lean | 12 ++++++---- .../RingTheory/UniqueFactorizationDomain.lean | 23 ++++++++++++++----- 2 files changed, 25 insertions(+), 10 deletions(-) diff --git a/Mathlib/RingTheory/Radical.lean b/Mathlib/RingTheory/Radical.lean index c86cee4872a46..0c28c161cdf4f 100644 --- a/Mathlib/RingTheory/Radical.lean +++ b/Mathlib/RingTheory/Radical.lean @@ -46,6 +46,12 @@ variable {M : Type*} [CancelCommMonoidWithZero M] [NormalizationMonoid M] def primeFactors (a : M) : Finset M := (normalizedFactors a).toFinset +theorem _root_.Associated.primeFactors_eq {a b : M} (h : Associated a b) : + primeFactors a = primeFactors b := by + unfold primeFactors + rw [h.normalizedFactors_eq] + + /-- Radical of an element `a` in a unique factorization monoid is the product of the prime factors of `a`. @@ -62,10 +68,8 @@ theorem radical_one_eq : radical (1 : M) = 1 := by rw [radical, primeFactors, normalizedFactors_one, Multiset.toFinset_zero, Finset.prod_empty] theorem radical_eq_of_associated {a b : M} (h : Associated a b) : radical a = radical b := by - rcases iff_iff_and_or_not_and_not.mp h.eq_zero_iff with (⟨rfl, rfl⟩ | ⟨ha, hb⟩) - · rfl - · simp_rw [radical, primeFactors] - rw [(associated_iff_normalizedFactors_eq_normalizedFactors ha hb).mp h] + unfold radical + rw [h.primeFactors_eq] theorem radical_unit_eq_one {a : M} (h : IsUnit a) : radical a = 1 := (radical_eq_of_associated (associated_one_iff_isUnit.mpr h)).trans radical_one_eq diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain.lean b/Mathlib/RingTheory/UniqueFactorizationDomain.lean index e144f86f1433b..dde98af42a008 100644 --- a/Mathlib/RingTheory/UniqueFactorizationDomain.lean +++ b/Mathlib/RingTheory/UniqueFactorizationDomain.lean @@ -554,6 +554,13 @@ theorem factors_pow_count_prod [DecidableEq α] {x : α} (hx : x ≠ 0) : _ = prod (factors x) := by rw [toFinset_sum_count_nsmul_eq (factors x)] _ ~ᵤ x := factors_prod hx +theorem factors_rel_of_associated {a b : α} (h : Associated a b) : + Multiset.Rel Associated (factors a) (factors b) := by + rcases iff_iff_and_or_not_and_not.mp h.eq_zero_iff with (⟨rfl, rfl⟩ | ⟨ha, hb⟩) + · simp + · refine factors_unique irreducible_of_factor irreducible_of_factor ?_ + exact ((factors_prod ha).trans h).trans (factors_prod hb).symm + end UniqueFactorizationMonoid namespace UniqueFactorizationMonoid @@ -721,13 +728,17 @@ theorem dvd_iff_normalizedFactors_le_normalizedFactors {x y : α} (hx : x ≠ 0) (normalizedFactors_prod hy).dvd_iff_dvd_right] apply Multiset.prod_dvd_prod_of_le +theorem _root_.Associated.normalizedFactors_eq {a b : α} (h : Associated a b) : + normalizedFactors a = normalizedFactors b := by + unfold normalizedFactors + have h' : ⇑(normalize (α := α)) = Associates.out ∘ Associates.mk := funext Associates.out_mk + rw [h', ← Multiset.map_map, ← Multiset.map_map, + Associates.rel_associated_iff_map_eq_map.mp (factors_rel_of_associated h)] + theorem associated_iff_normalizedFactors_eq_normalizedFactors {x y : α} (hx : x ≠ 0) (hy : y ≠ 0) : - x ~ᵤ y ↔ normalizedFactors x = normalizedFactors y := by - refine - ⟨fun h => ?_, fun h => - (normalizedFactors_prod hx).symm.trans (_root_.trans (by rw [h]) (normalizedFactors_prod hy))⟩ - apply le_antisymm <;> rw [← dvd_iff_normalizedFactors_le_normalizedFactors] - all_goals simp [*, h.dvd, h.symm.dvd] + x ~ᵤ y ↔ normalizedFactors x = normalizedFactors y := + ⟨Associated.normalizedFactors_eq, fun h => + (normalizedFactors_prod hx).symm.trans (_root_.trans (by rw [h]) (normalizedFactors_prod hy))⟩ theorem normalizedFactors_of_irreducible_pow {p : α} (hp : Irreducible p) (k : ℕ) : normalizedFactors (p ^ k) = Multiset.replicate k (normalize p) := by From 4092287b49c452c0efcf2dee7d9d5e5db095768d Mon Sep 17 00:00:00 2001 From: Etienne Date: Thu, 17 Oct 2024 19:04:10 +0000 Subject: [PATCH 108/505] feat: sufficient condition for a set of linear forms to span the dual space (#17548) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Consider a reflexive module and a set `s` of linear forms. If for any `z ≠ 0` there exists `f ∈ s` such that `f z ≠ 0`, then `s` spans the whole dual space. --- Mathlib/LinearAlgebra/Dual.lean | 14 ++++++++++++++ 1 file changed, 14 insertions(+) diff --git a/Mathlib/LinearAlgebra/Dual.lean b/Mathlib/LinearAlgebra/Dual.lean index 0dfa5a4d643f0..23b6f1f84347a 100644 --- a/Mathlib/LinearAlgebra/Dual.lean +++ b/Mathlib/LinearAlgebra/Dual.lean @@ -690,6 +690,20 @@ theorem exists_dual_map_eq_bot_of_lt_top (hp : p < ⊤) (hp' : Free R (M ⧸ p)) obtain ⟨f, hf, hf'⟩ := p.exists_dual_map_eq_bot_of_nmem hx hp' exact ⟨f, by aesop, hf'⟩ +/-- Consider a reflexive module and a set `s` of linear forms. If for any `z ≠ 0` there exists +`f ∈ s` such that `f z ≠ 0`, then `s` spans the whole dual space. -/ +theorem span_eq_top_of_ne_zero [IsReflexive R M] + {s : Set (M →ₗ[R] R)} [Free R ((M →ₗ[R] R) ⧸ (span R s))] + (h : ∀ z ≠ 0, ∃ f ∈ s, f z ≠ 0) : span R s = ⊤ := by + by_contra! hn + obtain ⟨φ, φne, hφ⟩ := exists_dual_map_eq_bot_of_lt_top hn.lt_top inferInstance + let φs := (evalEquiv R M).symm φ + have this f (hf : f ∈ s) : f φs = 0 := by + rw [← mem_bot R, ← hφ, mem_map] + exact ⟨f, subset_span hf, (apply_evalEquiv_symm_apply R M f φ).symm⟩ + obtain ⟨x, xs, hx⟩ := h φs (by simp [φne, φs]) + exact hx <| this x xs + variable {ι 𝕜 E : Type*} [Field 𝕜] [AddCommGroup E] [Module 𝕜 E] open LinearMap Set FiniteDimensional From b2dddd4c7b46f8feff35592d99c9e4ce57a3832f Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Thu, 17 Oct 2024 19:56:26 +0000 Subject: [PATCH 109/505] feat: sum of iterated derivatives (#16886) This introduces `Polynomial.sumIDeriv`, the sum of the iterated derivatives of a polynomial, to be used in particular in the proof of the Lindemann-Weierstrass theorem (see #6718). Included are: * `Polynomial.sumIDeriv_apply`, `Polynomial.sumIDeriv_apply_of_lt`, `Polynomial.sumIDeriv_apply_of_le`: `Polynomial.sumIDeriv` expressed as a sum * `Polynomial.sumIDeriv_C`, `Polynomial.sumIDeriv_X`: `Polynomial.sumIDeriv` applied to simple polynomials * `Polynomial.sumIDeriv_map`: `Polynomial.sumIDeriv` commutes with `Polynomial.map` * `Polynomial.sumIDeriv_derivative`: `Polynomial.sumIDeriv` commutes with `Polynomial.derivative` * `Polynomial.sumIDeriv_eq_self_add`: `sumIDeriv p = p + sumIDeriv (derivative p)` * `Polynomial.exists_iterate_derivative_eq_factorial_smul`: the `k`'th iterated derivative of a polynomial has a common factor `k!` * `Polynomial.aeval_iterate_derivative_of_lt`, `Polynomial.aeval_iterate_derivative_self`, `Polynomial.aeval_iterate_derivative_of_ge`: applying `Polynomial.aeval` to iterated derivatives * `Polynomial.aeval_sumIDeriv`, `Polynomial.aeval_sumIDeriv_of_pos`: applying `Polynomial.aeval` to `Polynomial.sumIDeriv` Co-authored-by: FR-vdash-bot Co-authored-by: FR --- Mathlib.lean | 1 + Mathlib/Algebra/Polynomial/Derivative.lean | 64 +++++ .../Polynomial/SumIteratedDerivative.lean | 245 ++++++++++++++++++ 3 files changed, 310 insertions(+) create mode 100644 Mathlib/Algebra/Polynomial/SumIteratedDerivative.lean diff --git a/Mathlib.lean b/Mathlib.lean index b43069a00a054..91da140ade63c 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -744,6 +744,7 @@ import Mathlib.Algebra.Polynomial.Roots import Mathlib.Algebra.Polynomial.Smeval import Mathlib.Algebra.Polynomial.SpecificDegree import Mathlib.Algebra.Polynomial.Splits +import Mathlib.Algebra.Polynomial.SumIteratedDerivative import Mathlib.Algebra.Polynomial.Taylor import Mathlib.Algebra.Polynomial.UnitTrinomial import Mathlib.Algebra.QuadraticDiscriminant diff --git a/Mathlib/Algebra/Polynomial/Derivative.lean b/Mathlib/Algebra/Polynomial/Derivative.lean index 9f22f0fa03689..9108354da435c 100644 --- a/Mathlib/Algebra/Polynomial/Derivative.lean +++ b/Mathlib/Algebra/Polynomial/Derivative.lean @@ -12,6 +12,7 @@ import Mathlib.GroupTheory.GroupAction.Ring ## Main definitions * `Polynomial.derivative`: The formal derivative of polynomials, expressed as a linear map. + * `Polynomial.derivativeFinsupp`: Iterated derivatives as a finite support function. -/ @@ -22,6 +23,8 @@ open Finset open Polynomial +open scoped Nat + namespace Polynomial universe u v w y z @@ -330,6 +333,21 @@ theorem coeff_iterate_derivative {k} (p : R[X]) (m : ℕ) : _ = Nat.descFactorial (m + k.succ) k.succ • p.coeff (m + k.succ) := by rw [Nat.succ_add_eq_add_succ] +theorem iterate_derivative_eq_sum (p : R[X]) (k : ℕ) : + derivative^[k] p = + ∑ x ∈ (derivative^[k] p).support, C ((x + k).descFactorial k • p.coeff (x + k)) * X ^ x := by + conv_lhs => rw [(derivative^[k] p).as_sum_support_C_mul_X_pow] + refine sum_congr rfl fun i _ ↦ ?_ + rw [coeff_iterate_derivative, Nat.descFactorial_eq_factorial_mul_choose] + +theorem iterate_derivative_eq_factorial_smul_sum (p : R[X]) (k : ℕ) : + derivative^[k] p = k ! • + ∑ x ∈ (derivative^[k] p).support, C ((x + k).choose k • p.coeff (x + k)) * X ^ x := by + conv_lhs => rw [iterate_derivative_eq_sum] + rw [smul_sum] + refine sum_congr rfl fun i _ ↦ ?_ + rw [← smul_mul_assoc, smul_C, smul_smul, Nat.descFactorial_eq_factorial_mul_choose] + theorem iterate_derivative_mul {n} (p q : R[X]) : derivative^[n] (p * q) = ∑ k ∈ range n.succ, (n.choose k • (derivative^[n - k] p * derivative^[k] q)) := by @@ -374,6 +392,52 @@ theorem iterate_derivative_mul {n} (p q : R[X]) : omega · rw [Nat.choose_zero_right, tsub_zero] +/-- +Iterated derivatives as a finite support function. +-/ +@[simps! apply_toFun] +noncomputable def derivativeFinsupp : R[X] →ₗ[R] ℕ →₀ R[X] where + toFun p := .onFinset (range (p.natDegree + 1)) (derivative^[·] p) fun i ↦ by + contrapose; simp_all [iterate_derivative_eq_zero, Nat.succ_le] + map_add' _ _ := by ext; simp + map_smul' _ _ := by ext; simp + +@[simp] +theorem support_derivativeFinsupp_subset_range {p : R[X]} {n : ℕ} (h : p.natDegree < n) : + (derivativeFinsupp p).support ⊆ range n := by + dsimp [derivativeFinsupp] + exact Finsupp.support_onFinset_subset.trans (Finset.range_subset.mpr h) + +@[simp] +theorem derivativeFinsupp_C (r : R) : derivativeFinsupp (C r : R[X]) = .single 0 (C r) := by + ext i : 1 + match i with + | 0 => simp + | i + 1 => simp + +@[simp] +theorem derivativeFinsupp_one : derivativeFinsupp (1 : R[X]) = .single 0 1 := by + simpa using derivativeFinsupp_C (1 : R) + +@[simp] +theorem derivativeFinsupp_X : derivativeFinsupp (X : R[X]) = .single 0 X + .single 1 1 := by + ext i : 1 + match i with + | 0 => simp + | 1 => simp + | (n + 2) => simp + +theorem derivativeFinsupp_map [Semiring S] (p : R[X]) (f : R →+* S) : + derivativeFinsupp (p.map f) = (derivativeFinsupp p).mapRange (·.map f) (by simp) := by + ext i : 1 + simp + +theorem derivativeFinsupp_derivative (p : R[X]) : + derivativeFinsupp (derivative p) = + (derivativeFinsupp p).comapDomain Nat.succ Nat.succ_injective.injOn := by + ext i : 1 + simp + end Semiring section CommSemiring diff --git a/Mathlib/Algebra/Polynomial/SumIteratedDerivative.lean b/Mathlib/Algebra/Polynomial/SumIteratedDerivative.lean new file mode 100644 index 0000000000000..b8b64c16ec9bf --- /dev/null +++ b/Mathlib/Algebra/Polynomial/SumIteratedDerivative.lean @@ -0,0 +1,245 @@ +/- +Copyright (c) 2022 Yuyang Zhao. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yuyang Zhao +-/ +import Mathlib.Algebra.Polynomial.AlgebraMap +import Mathlib.Algebra.Polynomial.BigOperators +import Mathlib.Algebra.Polynomial.Degree.Lemmas +import Mathlib.Algebra.Polynomial.Derivative +import Mathlib.Algebra.Polynomial.RingDivision + +/-! +# Sum of iterated derivatives + +This file introduces `Polynomial.sumIDeriv`, the sum of the iterated derivatives of a polynomial, +as a linear map. This is used in particular in the proof of the Lindemann-Weierstrass theorem +(see #6718). + +## Main results + +* `Polynomial.sumIDeriv`: Sum of iterated derivatives of a polynomial, as a linear map +* `Polynomial.sumIDeriv_apply`, `Polynomial.sumIDeriv_apply_of_lt`, + `Polynomial.sumIDeriv_apply_of_le`: `Polynomial.sumIDeriv` expressed as a sum +* `Polynomial.sumIDeriv_C`, `Polynomial.sumIDeriv_X`: `Polynomial.sumIDeriv` applied to simple + polynomials +* `Polynomial.sumIDeriv_map`: `Polynomial.sumIDeriv` commutes with `Polynomial.map` +* `Polynomial.sumIDeriv_derivative`: `Polynomial.sumIDeriv` commutes with `Polynomial.derivative` +* `Polynomial.sumIDeriv_eq_self_add`: `sumIDeriv p = p + sumIDeriv (derivative p)` +* `Polynomial.exists_iterate_derivative_eq_factorial_smul`: the `k`'th iterated derivative of a + polynomial has a common factor `k!` +* `Polynomial.aeval_iterate_derivative_of_lt`, `Polynomial.aeval_iterate_derivative_self`, + `Polynomial.aeval_iterate_derivative_of_ge`: applying `Polynomial.aeval` to iterated derivatives +* `Polynomial.aeval_sumIDeriv`, `Polynomial.aeval_sumIDeriv_of_pos`: applying `Polynomial.aeval` to + `Polynomial.sumIDeriv` + +-/ + +open Finset +open scoped Nat + +namespace Polynomial + +variable {R S : Type*} + +section Semiring + +variable [Semiring R] [Semiring S] + +/-- +Sum of iterated derivatives of a polynomial, as a linear map + +This definition does not allow different weights for the derivatives. It is likely that it could be +extended to allow them, but this was not needed for the initial use case (the integration by part +of the integral $I_i$ in the +[Lindemann-Weierstrass](https://en.wikipedia.org/wiki/Lindemann%E2%80%93Weierstrass_theorem) +theorem). +-/ +noncomputable def sumIDeriv : R[X] →ₗ[R] R[X] := + Finsupp.lsum ℕ (fun _ ↦ LinearMap.id) ∘ₗ derivativeFinsupp + +theorem sumIDeriv_apply (p : R[X]) : + sumIDeriv p = ∑ i ∈ range (p.natDegree + 1), derivative^[i] p := by + dsimp [sumIDeriv] + exact Finsupp.sum_of_support_subset _ (by simp) _ (by simp) + +theorem sumIDeriv_apply_of_lt {p : R[X]} {n : ℕ} (hn : p.natDegree < n) : + sumIDeriv p = ∑ i ∈ range n, derivative^[i] p := by + dsimp [sumIDeriv] + exact Finsupp.sum_of_support_subset _ (by simp [hn]) _ (by simp) + +theorem sumIDeriv_apply_of_le {p : R[X]} {n : ℕ} (hn : p.natDegree ≤ n) : + sumIDeriv p = ∑ i ∈ range (n + 1), derivative^[i] p := by + dsimp [sumIDeriv] + exact Finsupp.sum_of_support_subset _ (by simp [Nat.lt_succ, hn]) _ (by simp) + +@[simp] +theorem sumIDeriv_C (a : R) : sumIDeriv (C a) = C a := by + rw [sumIDeriv_apply, natDegree_C, zero_add, sum_range_one, Function.iterate_zero_apply] + +@[simp] +theorem sumIDeriv_X : sumIDeriv X = X + C 1 := by + rw [sumIDeriv_apply, natDegree_X, sum_range_succ, sum_range_one, Function.iterate_zero_apply, + Function.iterate_one, derivative_X, eq_natCast, Nat.cast_one] + +@[simp] +theorem sumIDeriv_map (p : R[X]) (f : R →+* S) : + sumIDeriv (p.map f) = (sumIDeriv p).map f := by + let n := max (p.map f).natDegree p.natDegree + rw [sumIDeriv_apply_of_le (le_max_left _ _ : _ ≤ n)] + rw [sumIDeriv_apply_of_le (le_max_right _ _ : _ ≤ n)] + simp_rw [Polynomial.map_sum, iterate_derivative_map p f] + +theorem sumIDeriv_derivative (p : R[X]) : sumIDeriv (derivative p) = derivative (sumIDeriv p) := by + rw [sumIDeriv_apply_of_le ((natDegree_derivative_le p).trans tsub_le_self), sumIDeriv_apply, + derivative_sum] + simp_rw [← Function.iterate_succ_apply, Function.iterate_succ_apply'] + +theorem sumIDeriv_eq_self_add (p : R[X]) : sumIDeriv p = p + sumIDeriv (derivative p) := by + rw [sumIDeriv_derivative, sumIDeriv_apply, derivative_sum, sum_range_succ', sum_range_succ, + add_comm, ← add_zero (Finset.sum _ _)] + simp_rw [← Function.iterate_succ_apply' derivative, Nat.succ_eq_add_one, + Function.iterate_zero_apply, iterate_derivative_eq_zero (Nat.lt_succ_self _)] + +theorem exists_iterate_derivative_eq_factorial_smul (p : R[X]) (k : ℕ) : + ∃ gp : R[X], gp.natDegree ≤ p.natDegree - k ∧ derivative^[k] p = k ! • gp := by + refine ⟨_, (natDegree_sum_le _ _).trans ?_, iterate_derivative_eq_factorial_smul_sum p k⟩ + rw [fold_max_le] + refine ⟨Nat.zero_le _, fun i hi => ?_⟩ + dsimp only [Function.comp] + exact (natDegree_C_mul_le _ _).trans <| (natDegree_X_pow_le _).trans <| + (le_natDegree_of_mem_supp _ hi).trans <| natDegree_iterate_derivative _ _ + +end Semiring + +section CommSemiring + +variable [CommSemiring R] {A : Type*} [CommRing A] [Algebra R A] + +theorem aeval_iterate_derivative_of_lt (p : R[X]) (q : ℕ) (r : A) {p' : A[X]} + (hp : p.map (algebraMap R A) = (X - C r) ^ q * p') {k : ℕ} (hk : k < q) : + aeval r (derivative^[k] p) = 0 := by + have h (x) : (X - C r) ^ (q - (k - x)) = (X - C r) ^ 1 * (X - C r) ^ (q - (k - x) - 1) := by + rw [← pow_add, add_tsub_cancel_of_le] + rw [Nat.lt_iff_add_one_le] at hk + exact (le_tsub_of_add_le_left hk).trans (tsub_le_tsub_left (tsub_le_self : _ ≤ k) _) + rw [aeval_def, eval₂_eq_eval_map, ← iterate_derivative_map] + simp_rw [hp, iterate_derivative_mul, iterate_derivative_X_sub_pow, ← smul_mul_assoc, smul_smul, + h, ← mul_smul_comm, mul_assoc, ← mul_sum, eval_mul, pow_one, eval_sub, eval_X, eval_C, sub_self, + zero_mul] + +theorem aeval_iterate_derivative_self (p : R[X]) (q : ℕ) (r : A) {p' : A[X]} + (hp : p.map (algebraMap R A) = (X - C r) ^ q * p') : + aeval r (derivative^[q] p) = q ! • p'.eval r := by + have h (x) (h : 1 ≤ x) (h' : x ≤ q) : + (X - C r) ^ (q - (q - x)) = (X - C r) ^ 1 * (X - C r) ^ (q - (q - x) - 1) := by + rw [← pow_add, add_tsub_cancel_of_le] + rwa [tsub_tsub_cancel_of_le h'] + rw [aeval_def, eval₂_eq_eval_map, ← iterate_derivative_map] + simp_rw [hp, iterate_derivative_mul, iterate_derivative_X_sub_pow, ← smul_mul_assoc, smul_smul] + rw [sum_range_succ', Nat.choose_zero_right, one_mul, tsub_zero, Nat.descFactorial_self, tsub_self, + pow_zero, smul_mul_assoc, one_mul, Function.iterate_zero_apply, eval_add, eval_smul] + convert zero_add _ + rw [eval_finset_sum] + apply sum_eq_zero + intro x hx + rw [h (x + 1) le_add_self (Nat.add_one_le_iff.mpr (mem_range.mp hx)), pow_one, + eval_mul, eval_smul, eval_mul, eval_sub, eval_X, eval_C, sub_self, zero_mul, + smul_zero, zero_mul] + +variable (A) + +theorem aeval_iterate_derivative_of_ge (p : R[X]) (q : ℕ) {k : ℕ} (hk : q ≤ k) : + ∃ gp : R[X], gp.natDegree ≤ p.natDegree - k ∧ + ∀ r : A, aeval r (derivative^[k] p) = q ! • aeval r gp := by + obtain ⟨p', p'_le, hp'⟩ := exists_iterate_derivative_eq_factorial_smul p k + obtain ⟨k, rfl⟩ := Nat.exists_eq_add_of_le hk + refine ⟨((q + k).descFactorial k : R[X]) * p', (natDegree_C_mul_le _ _).trans p'_le, fun r => ?_⟩ + simp_rw [hp', nsmul_eq_mul, map_mul, map_natCast, ← mul_assoc, ← Nat.cast_mul, + Nat.add_descFactorial_eq_ascFactorial, Nat.factorial_mul_ascFactorial] + +theorem aeval_sumIDeriv (p : R[X]) (q : ℕ) : + ∃ gp : R[X], gp.natDegree ≤ p.natDegree - q ∧ + ∀ (r : A) {p' : A[X]}, p.map (algebraMap R A) = (X - C r) ^ q * p' → + aeval r (sumIDeriv p) = q ! • aeval r gp := by + have h (k) : + ∃ gp : R[X], gp.natDegree ≤ p.natDegree - q ∧ + ∀ (r : A) {p' : A[X]}, p.map (algebraMap R A) = (X - C r) ^ q * p' → + aeval r (derivative^[k] p) = q ! • aeval r gp := by + cases lt_or_ge k q with + | inl hk => + use 0 + rw [natDegree_zero] + use Nat.zero_le _ + intro r p' hp + rw [map_zero, smul_zero, aeval_iterate_derivative_of_lt p q r hp hk] + | inr hk => + obtain ⟨gp, gp_le, h⟩ := aeval_iterate_derivative_of_ge A p q hk + exact ⟨gp, gp_le.trans (tsub_le_tsub_left hk _), fun r p' _ => h r⟩ + choose c h using h + choose c_le hc using h + refine ⟨(range (p.natDegree + 1)).sum c, ?_, ?_⟩ + · refine (natDegree_sum_le _ _).trans ?_ + rw [fold_max_le] + exact ⟨Nat.zero_le _, fun i _ => c_le i⟩ + intro r p' hp + rw [sumIDeriv_apply, map_sum]; simp_rw [hc _ r hp, map_sum, smul_sum] + +theorem aeval_sumIDeriv_of_pos [Nontrivial A] [NoZeroDivisors A] (p : R[X]) {q : ℕ} (hq : 0 < q) : + ∃ gp : R[X], gp.natDegree ≤ p.natDegree - q ∧ + ∀ (inj_amap : Function.Injective (algebraMap R A)) (r : A) {p' : A[X]}, + p.map (algebraMap R A) = (X - C r) ^ (q - 1) * p' → + aeval r (sumIDeriv p) = (q - 1)! • p'.eval r + q ! • aeval r gp := by + rcases eq_or_ne p 0 with (rfl | p0) + · use 0 + rw [natDegree_zero] + use Nat.zero_le _ + intro _ r p' hp + rw [map_zero, map_zero, smul_zero, add_zero] + rw [Polynomial.map_zero] at hp + replace hp := (mul_eq_zero.mp hp.symm).resolve_left ?_ + rw [hp, eval_zero, smul_zero] + exact fun h => X_sub_C_ne_zero r (pow_eq_zero h) + let c k := if hk : q ≤ k then (aeval_iterate_derivative_of_ge A p q hk).choose else 0 + have c_le (k) : (c k).natDegree ≤ p.natDegree - k := by + dsimp only [c] + split_ifs with h + · exact (aeval_iterate_derivative_of_ge A p q h).choose_spec.1 + · rw [natDegree_zero]; exact Nat.zero_le _ + have hc (k) (hk : q ≤ k) : ∀ (r : A), aeval r (derivative^[k] p) = q ! • aeval r (c k) := by + simp_rw [c, dif_pos hk] + exact (aeval_iterate_derivative_of_ge A p q hk).choose_spec.2 + refine ⟨∑ x ∈ Ico q (p.natDegree + 1), c x, ?_, ?_⟩ + · refine (natDegree_sum_le _ _).trans ?_ + rw [fold_max_le] + exact ⟨Nat.zero_le _, fun i hi => (c_le i).trans (tsub_le_tsub_left (mem_Ico.mp hi).1 _)⟩ + intro inj_amap r p' hp + have : range (p.natDegree + 1) = range q ∪ Ico q (p.natDegree + 1) := by + rw [range_eq_Ico, Ico_union_Ico_eq_Ico hq.le] + have h := natDegree_map_le (algebraMap R A) p + rw [congr_arg natDegree hp, natDegree_mul, natDegree_pow, natDegree_X_sub_C, mul_one, + ← Nat.sub_add_comm (Nat.one_le_of_lt hq), tsub_le_iff_right] at h + exact le_of_add_le_left h + · exact pow_ne_zero _ (X_sub_C_ne_zero r) + · rintro rfl + rw [mul_zero, Polynomial.map_eq_zero_iff inj_amap] at hp + exact p0 hp + rw [← zero_add ((q - 1)! • p'.eval r)] + rw [sumIDeriv_apply, map_sum, map_sum, this] + have : range q = range (q - 1 + 1) := by rw [tsub_add_cancel_of_le (Nat.one_le_of_lt hq)] + rw [sum_union, this, sum_range_succ] + congr 2 + · apply sum_eq_zero + exact fun x hx => aeval_iterate_derivative_of_lt p _ r hp (mem_range.mp hx) + · rw [← aeval_iterate_derivative_self _ _ _ hp] + · rw [smul_sum, sum_congr rfl] + intro k hk + exact hc k (mem_Ico.mp hk).1 r + · rw [range_eq_Ico, disjoint_iff_inter_eq_empty, eq_empty_iff_forall_not_mem] + intro x hx + rw [mem_inter, mem_Ico, mem_Ico] at hx + exact hx.1.2.not_le hx.2.1 + +end CommSemiring + +end Polynomial From bdecdcb60f733cb48871ea7bf782a46027f518d7 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Thu, 17 Oct 2024 19:56:27 +0000 Subject: [PATCH 110/505] chore: fix outdated `ext` porting notes (#17810) I went through all porting notes mentioning `ext` and fixed / removed those that no longer apply. --- Mathlib/Algebra/Polynomial/Monomial.lean | 9 +++------ Mathlib/AlgebraicGeometry/Limits.lean | 1 - Mathlib/CategoryTheory/EpiMono.lean | 2 -- Mathlib/Data/Set/Image.lean | 2 +- Mathlib/LinearAlgebra/Finsupp.lean | 6 ++---- Mathlib/LinearAlgebra/Matrix/Adjugate.lean | 3 +-- Mathlib/LinearAlgebra/Matrix/Charpoly/LinearMap.lean | 3 +-- Mathlib/NumberTheory/Zsqrtd/Basic.lean | 3 +-- Mathlib/RepresentationTheory/Invariants.lean | 2 +- 9 files changed, 10 insertions(+), 21 deletions(-) diff --git a/Mathlib/Algebra/Polynomial/Monomial.lean b/Mathlib/Algebra/Polynomial/Monomial.lean index 29b15cda6f45e..535103330d899 100644 --- a/Mathlib/Algebra/Polynomial/Monomial.lean +++ b/Mathlib/Algebra/Polynomial/Monomial.lean @@ -56,12 +56,9 @@ theorem ringHom_ext {S} [Semiring S] {f g : R[X] →+* S} (h₁ : ∀ a, f (C a) set f' := f.comp (toFinsuppIso R).symm.toRingHom with hf' set g' := g.comp (toFinsuppIso R).symm.toRingHom with hg' have A : f' = g' := by - -- Porting note: Was `ext; simp [..]; simpa [..] using h₂`. - ext : 1 - · ext - simp [f', g', h₁, RingEquiv.toRingHom_eq_coe] - · refine MonoidHom.ext_mnat ?_ - simpa [RingEquiv.toRingHom_eq_coe] using h₂ + ext + simp [f', g', h₁, RingEquiv.toRingHom_eq_coe] + simpa using h₂ have B : f = f'.comp (toFinsuppIso R) := by rw [hf', RingHom.comp_assoc] ext x diff --git a/Mathlib/AlgebraicGeometry/Limits.lean b/Mathlib/AlgebraicGeometry/Limits.lean index d5614273639d9..0622355eef447 100644 --- a/Mathlib/AlgebraicGeometry/Limits.lean +++ b/Mathlib/AlgebraicGeometry/Limits.lean @@ -57,7 +57,6 @@ def Scheme.emptyTo (X : Scheme.{u}) : ∅ ⟶ X := @[ext] theorem Scheme.empty_ext {X : Scheme.{u}} (f g : ∅ ⟶ X) : f = g := - -- Porting note (#11041): `ext` regression Scheme.Hom.ext' (Subsingleton.elim (α := ∅ ⟶ _) _ _) theorem Scheme.eq_emptyTo {X : Scheme.{u}} (f : ∅ ⟶ X) : f = Scheme.emptyTo X := diff --git a/Mathlib/CategoryTheory/EpiMono.lean b/Mathlib/CategoryTheory/EpiMono.lean index 13cfd3e4ad1b3..8c8f738cc3b9f 100644 --- a/Mathlib/CategoryTheory/EpiMono.lean +++ b/Mathlib/CategoryTheory/EpiMono.lean @@ -38,7 +38,6 @@ such that `f ≫ retraction f = 𝟙 X`. Every split monomorphism is a monomorphism. -/ /- Porting note(#5171): removed @[nolint has_nonempty_instance] -/ -/- Porting note: `@[ext]` used to accept lemmas like this. Now we add an aesop rule -/ @[ext, aesop apply safe (rule_sets := [CategoryTheory])] structure SplitMono {X Y : C} (f : X ⟶ Y) where /-- The map splitting `f` -/ @@ -64,7 +63,6 @@ such that `section_ f ≫ f = 𝟙 Y`. Every split epimorphism is an epimorphism. -/ /- Porting note(#5171): removed @[nolint has_nonempty_instance] -/ -/- Porting note: `@[ext]` used to accept lemmas like this. Now we add an aesop rule -/ @[ext, aesop apply safe (rule_sets := [CategoryTheory])] structure SplitEpi {X Y : C} (f : X ⟶ Y) where /-- The map splitting `f` -/ diff --git a/Mathlib/Data/Set/Image.lean b/Mathlib/Data/Set/Image.lean index 071d9c15c3f6c..00387c99f2202 100644 --- a/Mathlib/Data/Set/Image.lean +++ b/Mathlib/Data/Set/Image.lean @@ -950,7 +950,7 @@ theorem range_inclusion (h : s ⊆ t) : range (inclusion h) = { x : t | (x : α) -- When `f` is injective, see also `Equiv.ofInjective`. theorem leftInverse_rangeSplitting (f : α → β) : LeftInverse (rangeFactorization f) (rangeSplitting f) := fun x => by - apply Subtype.ext -- Porting note: why doesn't `ext` find this lemma? + ext simp only [rangeFactorization_coe] apply apply_rangeSplitting diff --git a/Mathlib/LinearAlgebra/Finsupp.lean b/Mathlib/LinearAlgebra/Finsupp.lean index 91e9feb4155da..5371c45091225 100644 --- a/Mathlib/LinearAlgebra/Finsupp.lean +++ b/Mathlib/LinearAlgebra/Finsupp.lean @@ -1331,8 +1331,7 @@ def splittingOfFinsuppSurjective (f : M →ₗ[R] α →₀ R) (s : Surjective f theorem splittingOfFinsuppSurjective_splits (f : M →ₗ[R] α →₀ R) (s : Surjective f) : f.comp (splittingOfFinsuppSurjective f s) = LinearMap.id := by - -- Porting note: `ext` can't find appropriate theorems. - refine lhom_ext' fun x => ext_ring <| Finsupp.ext fun y => ?_ + ext x dsimp [splittingOfFinsuppSurjective] congr rw [sum_single_index, one_smul] @@ -1357,8 +1356,7 @@ def splittingOfFunOnFintypeSurjective [Finite α] (f : M →ₗ[R] α → R) (s theorem splittingOfFunOnFintypeSurjective_splits [Finite α] (f : M →ₗ[R] α → R) (s : Surjective f) : f.comp (splittingOfFunOnFintypeSurjective f s) = LinearMap.id := by classical - -- Porting note: `ext` can't find appropriate theorems. - refine pi_ext' fun x => ext_ring <| funext fun y => ?_ + ext x y dsimp [splittingOfFunOnFintypeSurjective] rw [linearEquivFunOnFinite_symm_single, Finsupp.sum_single_index, one_smul, (s (Finsupp.single x 1)).choose_spec, Finsupp.single_eq_pi_single] diff --git a/Mathlib/LinearAlgebra/Matrix/Adjugate.lean b/Mathlib/LinearAlgebra/Matrix/Adjugate.lean index ef3a1b2ec6823..516ba73963c50 100644 --- a/Mathlib/LinearAlgebra/Matrix/Adjugate.lean +++ b/Mathlib/LinearAlgebra/Matrix/Adjugate.lean @@ -114,8 +114,7 @@ theorem cramer_row_self (i : n) (h : ∀ j, b j = A j i) : A.cramer b = Pi.singl @[simp] theorem cramer_one : cramer (1 : Matrix n n α) = 1 := by - -- Porting note: was `ext i j` - refine LinearMap.pi_ext' (fun (i : n) => LinearMap.ext_ring (funext (fun (j : n) => ?_))) + ext i j convert congr_fun (cramer_row_self (1 : Matrix n n α) (Pi.single i 1) i _) j · simp · intro j diff --git a/Mathlib/LinearAlgebra/Matrix/Charpoly/LinearMap.lean b/Mathlib/LinearAlgebra/Matrix/Charpoly/LinearMap.lean index e7d285c6ad0d2..a6138962526ff 100644 --- a/Mathlib/LinearAlgebra/Matrix/Charpoly/LinearMap.lean +++ b/Mathlib/LinearAlgebra/Matrix/Charpoly/LinearMap.lean @@ -93,8 +93,7 @@ theorem Matrix.represents_iff' {A : Matrix ι ι R} {f : Module.End R M} : have := LinearMap.congr_fun h (Pi.single i 1) rwa [PiToModule.fromEnd_apply_single_one, PiToModule.fromMatrix_apply_single_one] at this · intro h - -- Porting note: was `ext` - refine LinearMap.pi_ext' (fun i => LinearMap.ext_ring ?_) + ext simp_rw [LinearMap.comp_apply, LinearMap.coe_single, PiToModule.fromEnd_apply_single_one, PiToModule.fromMatrix_apply_single_one] apply h diff --git a/Mathlib/NumberTheory/Zsqrtd/Basic.lean b/Mathlib/NumberTheory/Zsqrtd/Basic.lean index 58b19978fb5b9..914914db4f6cb 100644 --- a/Mathlib/NumberTheory/Zsqrtd/Basic.lean +++ b/Mathlib/NumberTheory/Zsqrtd/Basic.lean @@ -931,8 +931,7 @@ def lift {d : ℤ} : { r : R // r * r = ↑d } ≃ (ℤ√d →+* R) where ext simp right_inv f := by - -- Porting note: was `ext` - refine hom_ext _ _ ?_ + ext simp /-- `lift r` is injective if `d` is non-square, and R has characteristic zero (that is, the map from diff --git a/Mathlib/RepresentationTheory/Invariants.lean b/Mathlib/RepresentationTheory/Invariants.lean index fa62df0a43948..898d8d3427513 100644 --- a/Mathlib/RepresentationTheory/Invariants.lean +++ b/Mathlib/RepresentationTheory/Invariants.lean @@ -134,7 +134,7 @@ def invariantsEquivRepHom (X Y : Rep k G) : (linHom X.ρ Y.ρ).invariants ≃ₗ map_add' _ _ := rfl map_smul' _ _ := rfl invFun f := ⟨f.hom, fun g => (mem_invariants_iff_comm _ g).2 (f.comm g)⟩ - left_inv _ := by apply Subtype.ext; ext; rfl -- Porting note: Added `apply Subtype.ext` + left_inv _ := by ext; rfl right_inv _ := by ext; rfl end Rep From e355e86a77c6deb2205be0bc9a4983e13b25eed4 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Thu, 17 Oct 2024 19:56:28 +0000 Subject: [PATCH 111/505] chore: fix "proof was" porting notes (#17836) Go through all porting notes that contain the string "proof was" and try to fix them. The following steps were used: * Golf the new proof. * If the current proof is nicer, delete the note. * Try to restore the old proof. * Classify some of the old failing proofs. --- .../Algebra/Category/ModuleCat/Subobject.lean | 4 +-- Mathlib/Algebra/Order/Group/Units.lean | 6 +--- Mathlib/Algebra/Order/Interval/Basic.lean | 7 ++--- Mathlib/Algebra/Star/Subalgebra.lean | 10 +++--- .../Analysis/Calculus/FDeriv/Measurable.lean | 1 - Mathlib/Data/List/OfFn.lean | 10 ++---- Mathlib/Data/Tree/Basic.lean | 4 +-- .../IsAlgClosed/AlgebraicClosure.lean | 31 +++++++------------ Mathlib/Geometry/Manifold/ChartedSpace.lean | 3 +- .../LinearAlgebra/CliffordAlgebra/Equivs.lean | 3 +- .../CliffordAlgebra/EvenEquiv.lean | 9 ++---- Mathlib/LinearAlgebra/LinearIndependent.lean | 4 +-- Mathlib/LinearAlgebra/SymplecticGroup.lean | 1 - Mathlib/MeasureTheory/Measure/Hausdorff.lean | 14 ++------- Mathlib/Order/Hom/Lattice.lean | 6 ---- Mathlib/RepresentationTheory/Rep.lean | 6 +--- Mathlib/RingTheory/AdjoinRoot.lean | 3 -- Mathlib/RingTheory/WittVector/IsPoly.lean | 4 --- Mathlib/Topology/Algebra/Monoid.lean | 2 -- Mathlib/Topology/Constructions.lean | 2 -- .../Topology/FiberBundle/Constructions.lean | 7 ++--- 21 files changed, 33 insertions(+), 104 deletions(-) diff --git a/Mathlib/Algebra/Category/ModuleCat/Subobject.lean b/Mathlib/Algebra/Category/ModuleCat/Subobject.lean index 134f3702fde18..a1508f23b2e54 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Subobject.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Subobject.lean @@ -82,7 +82,7 @@ noncomputable def toKernelSubobject {M N : ModuleCat.{v} R} {f : M ⟶ N} : @[simp] theorem toKernelSubobject_arrow {M N : ModuleCat R} {f : M ⟶ N} (x : LinearMap.ker f) : (kernelSubobject f).arrow (toKernelSubobject x) = x.1 := by - -- Porting note: The whole proof was just `simp [toKernelSubobject]`. + -- Porting note (#10959): the whole proof was just `simp [toKernelSubobject]`. suffices ((arrow ((kernelSubobject f))) ∘ (kernelSubobjectIso f ≪≫ kernelIsoKer f).inv) x = x by convert this rw [Iso.trans_inv, ← coe_comp, Category.assoc] @@ -105,7 +105,7 @@ theorem cokernel_π_imageSubobject_ext {L M N : ModuleCat.{v} R} (f : L ⟶ M) [ (g : (imageSubobject f : ModuleCat.{v} R) ⟶ N) [HasCokernel g] {x y : N} (l : L) (w : x = y + g (factorThruImageSubobject f l)) : cokernel.π g x = cokernel.π g y := by subst w - -- Porting note: The proof from here used to just be `simp`. + -- Porting note (#10959): The proof from here used to just be `simp`. simp only [map_add, add_right_eq_self] change ((cokernel.π g) ∘ (g) ∘ (factorThruImageSubobject f)) l = 0 rw [← coe_comp, ← coe_comp, Category.assoc] diff --git a/Mathlib/Algebra/Order/Group/Units.lean b/Mathlib/Algebra/Order/Group/Units.lean index af674406ea62a..c32449fad0ca7 100644 --- a/Mathlib/Algebra/Order/Group/Units.lean +++ b/Mathlib/Algebra/Order/Group/Units.lean @@ -20,11 +20,7 @@ variable {α : Type*} additive group."] instance Units.orderedCommGroup [OrderedCommMonoid α] : OrderedCommGroup αˣ := { Units.instPartialOrderUnits, Units.instCommGroupUnits with - mul_le_mul_left := fun _ _ h _ => (@mul_le_mul_left' α _ _ _ _ _ h _) } - --- Porting note: the mathlib3 proof was --- mul_le_mul_left := fun a b h c => (mul_le_mul_left' (h : (a : α) ≤ b) _ : (c : α) * a ≤ c * b) } --- see https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/elaboration.20failure.20in.20algebra.2Eorder.2Egroup.2Eunits + mul_le_mul_left := fun _ _ h _ => (mul_le_mul_left' (α := α) h _) } /-- The units of a linearly ordered commutative monoid form a linearly ordered commutative group. -/ @[to_additive "The units of a linearly ordered commutative additive monoid form a diff --git a/Mathlib/Algebra/Order/Interval/Basic.lean b/Mathlib/Algebra/Order/Interval/Basic.lean index dcd8a5d261e09..5fc5f067acadf 100644 --- a/Mathlib/Algebra/Order/Interval/Basic.lean +++ b/Mathlib/Algebra/Order/Interval/Basic.lean @@ -611,11 +611,8 @@ theorem length_sub_le : (s - t).length ≤ s.length + t.length := by simpa [sub_eq_add_neg] using length_add_le s (-t) theorem length_sum_le (f : ι → Interval α) (s : Finset ι) : - (∑ i ∈ s, f i).length ≤ ∑ i ∈ s, (f i).length := by - -- Porting note: Old proof was `:= Finset.le_sum_of_subadditive _ length_zero length_add_le _ _` - apply Finset.le_sum_of_subadditive - · exact length_zero - · exact length_add_le + (∑ i ∈ s, f i).length ≤ ∑ i ∈ s, (f i).length := + Finset.le_sum_of_subadditive _ length_zero length_add_le _ _ end Interval diff --git a/Mathlib/Algebra/Star/Subalgebra.lean b/Mathlib/Algebra/Star/Subalgebra.lean index 24e84fa477d77..a88d2390accc9 100644 --- a/Mathlib/Algebra/Star/Subalgebra.lean +++ b/Mathlib/Algebra/Star/Subalgebra.lean @@ -107,8 +107,7 @@ theorem toSubalgebra_le_iff {S₁ S₂ : StarSubalgebra R A} : equalities. -/ protected def copy (S : StarSubalgebra R A) (s : Set A) (hs : s = ↑S) : StarSubalgebra R A where toSubalgebra := Subalgebra.copy S.toSubalgebra s hs - star_mem' := @fun a ha => hs ▸ (S.star_mem' (by simpa [hs] using ha) : star a ∈ (S : Set A)) - -- Porting note: the old proof kept crashing Lean + star_mem' {a} ha := hs ▸ S.star_mem' (by simpa [hs] using ha) @[simp] theorem coe_copy (S : StarSubalgebra R A) (s : Set A) (hs : s = ↑S) : (S.copy s hs : Set A) = s := @@ -700,10 +699,9 @@ theorem ext_adjoin_singleton {a : A} [FunLike F (adjoin R ({a} : Set A)) B] variable [FunLike F A B] [AlgHomClass F R A B] [StarHomClass F A B] (f g : F) /-- The equalizer of two star `R`-algebra homomorphisms. -/ -def equalizer : StarSubalgebra R A := - { toSubalgebra := AlgHom.equalizer (f : A →ₐ[R] B) g - star_mem' := @fun a (ha : f a = g a) => by simpa only [← map_star] using congrArg star ha } --- Porting note: much like `StarSubalgebra.copy` the old proof was broken and hard to fix +def equalizer : StarSubalgebra R A where + toSubalgebra := AlgHom.equalizer (f : A →ₐ[R] B) g + star_mem' {a} (ha : f a = g a) := by simpa only [← map_star] using congrArg star ha @[simp] theorem mem_equalizer (x : A) : x ∈ StarAlgHom.equalizer f g ↔ f x = g x := diff --git a/Mathlib/Analysis/Calculus/FDeriv/Measurable.lean b/Mathlib/Analysis/Calculus/FDeriv/Measurable.lean index 28e1539f0e516..5801ab5775fa7 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Measurable.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Measurable.lean @@ -592,7 +592,6 @@ theorem D_subset_differentiable_set {K : Set F} (hK : IsComplete K) : _ ≤ ‖L e p q - L e p r‖ + ‖L e p r - L e' p' r‖ + ‖L e' p' r - L e' p' q'‖ := (le_trans (norm_add_le _ _) (add_le_add_right (norm_add_le _ _) _)) _ ≤ 4 * (1 / 2) ^ e + 4 * (1 / 2) ^ e + 4 * (1 / 2) ^ e := by gcongr - -- Porting note: proof was `by apply_rules [add_le_add]` _ = 12 * (1 / 2) ^ e := by ring /- For definiteness, use `L0 e = L e (n e) (n e)`, to have a single sequence. We claim that this diff --git a/Mathlib/Data/List/OfFn.lean b/Mathlib/Data/List/OfFn.lean index 9aa74790b948b..a3271be86ffd8 100644 --- a/Mathlib/Data/List/OfFn.lean +++ b/Mathlib/Data/List/OfFn.lean @@ -200,14 +200,8 @@ def ofFnRec {C : List α → Sort*} (h : ∀ (n) (f : Fin n → α), C (List.ofF @[simp] theorem ofFnRec_ofFn {C : List α → Sort*} (h : ∀ (n) (f : Fin n → α), C (List.ofFn f)) {n : ℕ} - (f : Fin n → α) : @ofFnRec _ C h (List.ofFn f) = h _ f := by - -- Porting note: Old proof was - -- equivSigmaTuple.rightInverse_symm.cast_eq (fun s => h s.1 s.2) ⟨n, f⟩ - have := (@equivSigmaTuple α).rightInverse_symm - dsimp [equivSigmaTuple] at this - have := this.cast_eq (fun s => h s.1 s.2) ⟨n, f⟩ - dsimp only at this - rw [ofFnRec, ← this] + (f : Fin n → α) : @ofFnRec _ C h (List.ofFn f) = h _ f := + equivSigmaTuple.rightInverse_symm.cast_eq (fun s => h s.1 s.2) ⟨n, f⟩ theorem exists_iff_exists_tuple {P : List α → Prop} : (∃ l : List α, P l) ↔ ∃ (n : _) (f : Fin n → α), P (List.ofFn f) := diff --git a/Mathlib/Data/Tree/Basic.lean b/Mathlib/Data/Tree/Basic.lean index 95ebf3246593c..df733700e6e4a 100644 --- a/Mathlib/Data/Tree/Basic.lean +++ b/Mathlib/Data/Tree/Basic.lean @@ -103,9 +103,7 @@ compile_inductive% Tree @[elab_as_elim] def unitRecOn {motive : Tree Unit → Sort*} (t : Tree Unit) (base : motive nil) (ind : ∀ x y, motive x → motive y → motive (x △ y)) : motive t := - -- Porting note: Old proof was `t.recOn base fun u => u.recOn ind` but - -- structure eta makes it unnecessary (https://github.com/leanprover/lean4/issues/777). - t.recOn base fun _u => ind + t.recOn base fun _u ↦ ind theorem left_node_right_eq_self : ∀ {x : Tree Unit} (_hx : x ≠ nil), x.left △ x.right = x | nil, h => by trivial diff --git a/Mathlib/FieldTheory/IsAlgClosed/AlgebraicClosure.lean b/Mathlib/FieldTheory/IsAlgClosed/AlgebraicClosure.lean index a00d2cba8705e..ea7fe25e09f58 100644 --- a/Mathlib/FieldTheory/IsAlgClosed/AlgebraicClosure.lean +++ b/Mathlib/FieldTheory/IsAlgClosed/AlgebraicClosure.lean @@ -164,12 +164,8 @@ instance Step.algebraSucc (n) : Algebra (Step k n) (Step k (n + 1)) := (toStepSucc k n).toAlgebra theorem toStepSucc.exists_root {n} {f : Polynomial (Step k n)} (hfm : f.Monic) - (hfi : Irreducible f) : ∃ x : Step k (n + 1), f.eval₂ (toStepSucc k n) x = 0 := by --- Porting note: original proof was `@AdjoinMonic.exists_root _ (Step.field k n) _ hfm hfi`, --- but it timeouts. - obtain ⟨x, hx⟩ := @AdjoinMonic.exists_root _ (Step.field k n) _ hfm hfi --- Porting note: using `hx` instead of `by apply hx` timeouts. - exact ⟨x, by apply hx⟩ + (hfi : Irreducible f) : ∃ x : Step k (n + 1), f.eval₂ (toStepSucc k n) x = 0 := + @AdjoinMonic.exists_root _ (Step.field k n) _ hfm hfi -- Porting note: the following two declarations were added during the port to be used in the -- definition of toStepOfLE @@ -186,29 +182,24 @@ private theorem toStepOfLE'.succ (m n : ℕ) (h : m ≤ n) : def toStepOfLE (m n : ℕ) (h : m ≤ n) : Step k m →+* Step k n where toFun := toStepOfLE' k m n h map_one' := by --- Porting note: original proof was `induction' h with n h ih; · exact Nat.leRecOn_self 1` --- `rw [Nat.leRecOn_succ h, ih, RingHom.map_one]` induction' h with a h ih · exact Nat.leRecOn_self 1 - · rw [toStepOfLE'.succ k m a h]; simp [ih] + · simp [toStepOfLE'.succ k m a h, ih] map_mul' x y := by --- Porting note: original proof was `induction' h with n h ih; · simp_rw [Nat.leRecOn_self]` --- `simp_rw [Nat.leRecOn_succ h, ih, RingHom.map_mul]` + simp only induction' h with a h ih - · dsimp [toStepOfLE']; simp_rw [Nat.leRecOn_self] - · simp_rw [toStepOfLE'.succ k m a h]; simp only at ih; simp [ih] --- Porting note: original proof was `induction' h with n h ih; · exact Nat.leRecOn_self 0` --- `rw [Nat.leRecOn_succ h, ih, RingHom.map_zero]` + · simp_rw [toStepOfLE', Nat.leRecOn_self] + · simp [toStepOfLE'.succ k m a h, ih] map_zero' := by + simp only induction' h with a h ih · exact Nat.leRecOn_self 0 - · simp_rw [toStepOfLE'.succ k m a h]; simp only at ih; simp [ih] + · simp [toStepOfLE'.succ k m a h, ih] map_add' x y := by --- Porting note: original proof was `induction' h with n h ih; · simp_rw [Nat.leRecOn_self]` --- `simp_rw [Nat.leRecOn_succ h, ih, RingHom.map_add]` + simp only induction' h with a h ih - · dsimp [toStepOfLE']; simp_rw [Nat.leRecOn_self] - · simp_rw [toStepOfLE'.succ k m a h]; simp only at ih; simp [ih] + · simp_rw [toStepOfLE', Nat.leRecOn_self] + · simp [toStepOfLE'.succ k m a h, ih] @[simp] theorem coe_toStepOfLE (m n : ℕ) (h : m ≤ n) : diff --git a/Mathlib/Geometry/Manifold/ChartedSpace.lean b/Mathlib/Geometry/Manifold/ChartedSpace.lean index 6a69ac58da981..5acd6a8cea119 100644 --- a/Mathlib/Geometry/Manifold/ChartedSpace.lean +++ b/Mathlib/Geometry/Manifold/ChartedSpace.lean @@ -758,9 +758,8 @@ section variable {ι : Type*} {Hi : ι → Type*} --- Porting note: Old proof was `Pi.inhabited _`. instance modelPiInhabited [∀ i, Inhabited (Hi i)] : Inhabited (ModelPi Hi) := - ⟨fun _ ↦ default⟩ + Pi.instInhabited instance [∀ i, TopologicalSpace (Hi i)] : TopologicalSpace (ModelPi Hi) := Pi.topologicalSpace diff --git a/Mathlib/LinearAlgebra/CliffordAlgebra/Equivs.lean b/Mathlib/LinearAlgebra/CliffordAlgebra/Equivs.lean index 156e71925668f..2285548a3e107 100644 --- a/Mathlib/LinearAlgebra/CliffordAlgebra/Equivs.lean +++ b/Mathlib/LinearAlgebra/CliffordAlgebra/Equivs.lean @@ -392,7 +392,6 @@ theorem equiv_ι (r : R) : CliffordAlgebraDualNumber.equiv (ι (R := R) _ r) = r @[simp] theorem equiv_symm_eps : CliffordAlgebraDualNumber.equiv.symm (eps : R[ε]) = ι (0 : QuadraticForm R R) 1 := - -- Porting note: Original proof was `DualNumber.lift_apply_eps _` - DualNumber.lift_apply_eps (R := R) (B := CliffordAlgebra (0 : QuadraticForm R R)) _ + DualNumber.lift_apply_eps _ end CliffordAlgebraDualNumber diff --git a/Mathlib/LinearAlgebra/CliffordAlgebra/EvenEquiv.lean b/Mathlib/LinearAlgebra/CliffordAlgebra/EvenEquiv.lean index 119f96a21ae0d..e9ab3a38b7895 100644 --- a/Mathlib/LinearAlgebra/CliffordAlgebra/EvenEquiv.lean +++ b/Mathlib/LinearAlgebra/CliffordAlgebra/EvenEquiv.lean @@ -152,13 +152,8 @@ def ofEven : CliffordAlgebra.even (Q' Q) →ₐ[R] CliffordAlgebra Q := by theorem ofEven_ι (x y : M × R) : ofEven Q ((even.ι (Q' Q)).bilin x y) = - (ι Q x.1 + algebraMap R _ x.2) * (ι Q y.1 - algebraMap R _ y.2) := by - -- Porting note: entire proof was the term-mode `even.lift_ι (Q' Q) _ x y` - unfold ofEven - lift_lets - intro f - -- TODO: replacing `?_` with `_` takes way longer? - exact @even.lift_ι R (M × R) _ _ _ (Q' Q) _ _ _ ⟨f, ?_, ?_⟩ x y + (ι Q x.1 + algebraMap R _ x.2) * (ι Q y.1 - algebraMap R _ y.2) := + even.lift_ι (Q' Q) _ x y theorem toEven_comp_ofEven : (toEven Q).comp (ofEven Q) = AlgHom.id R _ := even.algHom_ext (Q' Q) <| diff --git a/Mathlib/LinearAlgebra/LinearIndependent.lean b/Mathlib/LinearAlgebra/LinearIndependent.lean index ccd018f0d3ddf..3f46f76b260ab 100644 --- a/Mathlib/LinearAlgebra/LinearIndependent.lean +++ b/Mathlib/LinearAlgebra/LinearIndependent.lean @@ -931,9 +931,7 @@ theorem LinearIndependent.independent_span_singleton (hv : LinearIndependent R v obtain ⟨⟨r, rfl⟩, hm⟩ := hm suffices r = 0 by simp [this] apply linearIndependent_iff_not_smul_mem_span.mp hv i - -- Porting note: The original proof was using `convert hm`. - suffices v '' (univ \ {i}) = range fun j : { j // j ≠ i } => v j by - rwa [this] + convert hm ext simp diff --git a/Mathlib/LinearAlgebra/SymplecticGroup.lean b/Mathlib/LinearAlgebra/SymplecticGroup.lean index c638f98ae8bd9..55b9036e4c5a4 100644 --- a/Mathlib/LinearAlgebra/SymplecticGroup.lean +++ b/Mathlib/LinearAlgebra/SymplecticGroup.lean @@ -86,7 +86,6 @@ open Matrix theorem mem_iff {A : Matrix (l ⊕ l) (l ⊕ l) R} : A ∈ symplecticGroup l R ↔ A * J l R * Aᵀ = J l R := by simp [symplecticGroup] --- Porting note: Previous proof was `by infer_instance` instance coeMatrix : Coe (symplecticGroup l R) (Matrix (l ⊕ l) (l ⊕ l) R) := ⟨Subtype.val⟩ diff --git a/Mathlib/MeasureTheory/Measure/Hausdorff.lean b/Mathlib/MeasureTheory/Measure/Hausdorff.lean index d4023bb58018d..ce69db687e782 100644 --- a/Mathlib/MeasureTheory/Measure/Hausdorff.lean +++ b/Mathlib/MeasureTheory/Measure/Hausdorff.lean @@ -1026,18 +1026,8 @@ theorem hausdorffMeasure_smul_right_image [NormedAddCommGroup E] [NormedSpace -- break lineMap into pieces suffices μH[1] ((‖v‖ • ·) '' (LinearMap.toSpanSingleton ℝ E (‖v‖⁻¹ • v) '' s)) = ‖v‖₊ • μH[1] s by - -- Porting note: proof was shorter, could need some golf - simp only [hausdorffMeasure_real, nnreal_smul_coe_apply] - convert this - · simp only [image_smul, LinearMap.toSpanSingleton_apply, Set.image_image] - ext e - simp only [mem_image] - refine ⟨fun ⟨x, h⟩ => ⟨x, ?_⟩, fun ⟨x, h⟩ => ⟨x, ?_⟩⟩ - · rw [smul_comm (norm _), smul_comm (norm _), inv_smul_smul₀ hn] - exact h - · rw [smul_comm (norm _), smul_comm (norm _), inv_smul_smul₀ hn] at h - exact h - · exact hausdorffMeasure_real.symm + simpa only [Set.image_image, smul_comm (norm _), inv_smul_smul₀ hn, + LinearMap.toSpanSingleton_apply] using this have iso_smul : Isometry (LinearMap.toSpanSingleton ℝ E (‖v‖⁻¹ • v)) := by refine AddMonoidHomClass.isometry_of_norm _ fun x => (norm_smul _ _).trans ?_ rw [norm_smul, norm_inv, norm_norm, inv_mul_cancel₀ hn, mul_one, LinearMap.id_apply] diff --git a/Mathlib/Order/Hom/Lattice.lean b/Mathlib/Order/Hom/Lattice.lean index 94af3f7566504..4179afdc3722b 100644 --- a/Mathlib/Order/Hom/Lattice.lean +++ b/Mathlib/Order/Hom/Lattice.lean @@ -1567,7 +1567,6 @@ theorem withTop_id : (SupHom.id α).withTop = SupHom.id _ := DFunLike.coe_inject @[simp] theorem withTop_comp (f : SupHom β γ) (g : SupHom α β) : (f.comp g).withTop = f.withTop.comp g.withTop := --- Porting note: Proof was `DFunLike.coe_injective (Option.map_comp_map _ _).symm` DFunLike.coe_injective <| Eq.symm <| Option.map_comp_map _ _ /-- Adjoins a `⊥` to the domain and codomain of a `SupHom`. -/ @@ -1588,7 +1587,6 @@ theorem withBot_id : (SupHom.id α).withBot = SupBotHom.id _ := DFunLike.coe_inj @[simp] theorem withBot_comp (f : SupHom β γ) (g : SupHom α β) : (f.comp g).withBot = f.withBot.comp g.withBot := --- Porting note: Proof was `DFunLike.coe_injective (Option.map_comp_map _ _).symm` DFunLike.coe_injective <| Eq.symm <| Option.map_comp_map _ _ /-- Adjoins a `⊤` to the codomain of a `SupHom`. -/ @@ -1638,7 +1636,6 @@ theorem withTop_id : (InfHom.id α).withTop = InfTopHom.id _ := DFunLike.coe_inj @[simp] theorem withTop_comp (f : InfHom β γ) (g : InfHom α β) : (f.comp g).withTop = f.withTop.comp g.withTop := --- Porting note: Proof was `DFunLike.coe_injective (Option.map_comp_map _ _).symm` DFunLike.coe_injective <| Eq.symm <| Option.map_comp_map _ _ /-- Adjoins a `⊥` to the domain and codomain of an `InfHom`. -/ @@ -1658,7 +1655,6 @@ theorem withBot_id : (InfHom.id α).withBot = InfHom.id _ := DFunLike.coe_inject @[simp] theorem withBot_comp (f : InfHom β γ) (g : InfHom α β) : (f.comp g).withBot = f.withBot.comp g.withBot := --- Porting note: Proof was `DFunLike.coe_injective (Option.map_comp_map _ _).symm` DFunLike.coe_injective <| Eq.symm <| Option.map_comp_map _ _ /-- Adjoins a `⊤` to the codomain of an `InfHom`. -/ @@ -1708,7 +1704,6 @@ theorem withTop_id : (LatticeHom.id α).withTop = LatticeHom.id _ := @[simp] theorem withTop_comp (f : LatticeHom β γ) (g : LatticeHom α β) : (f.comp g).withTop = f.withTop.comp g.withTop := --- Porting note: Proof was `DFunLike.coe_injective (Option.map_comp_map _ _).symm` DFunLike.coe_injective <| Eq.symm <| Option.map_comp_map _ _ /-- Adjoins a `⊥` to the domain and codomain of a `LatticeHom`. -/ @@ -1729,7 +1724,6 @@ theorem withBot_id : (LatticeHom.id α).withBot = LatticeHom.id _ := @[simp] theorem withBot_comp (f : LatticeHom β γ) (g : LatticeHom α β) : (f.comp g).withBot = f.withBot.comp g.withBot := --- Porting note: Proof was `DFunLike.coe_injective (Option.map_comp_map _ _).symm` DFunLike.coe_injective <| Eq.symm <| Option.map_comp_map _ _ /-- Adjoins a `⊤` and `⊥` to the domain and codomain of a `LatticeHom`. -/ diff --git a/Mathlib/RepresentationTheory/Rep.lean b/Mathlib/RepresentationTheory/Rep.lean index 8ebf69973568b..e6d2b8e5cd51e 100644 --- a/Mathlib/RepresentationTheory/Rep.lean +++ b/Mathlib/RepresentationTheory/Rep.lean @@ -183,13 +183,9 @@ theorem linearization_μ_hom (X Y : Action (Type u) (MonCat.of G)) : @[simp] theorem linearization_μ_inv_hom (X Y : Action (Type u) (MonCat.of G)) : (inv ((linearization k G).μ X Y)).hom = (finsuppTensorFinsupp' k X.V Y.V).symm.toLinearMap := by --- Porting note (#11039): broken proof was -/- simp_rw [← Action.forget_map, Functor.map_inv, Action.forget_map, linearization_μ_hom] - apply IsIso.inv_eq_of_hom_inv_id _ - exact LinearMap.ext fun x => LinearEquiv.symm_apply_apply _ _-/ rw [← Action.forget_map, Functor.map_inv] apply IsIso.inv_eq_of_hom_inv_id - exact LinearMap.ext fun x => LinearEquiv.symm_apply_apply (finsuppTensorFinsupp' k X.V Y.V) x + exact LinearMap.ext fun x ↦ LinearEquiv.symm_apply_apply _ _ @[simp] theorem linearization_ε_hom : (linearization k G).ε.hom = Finsupp.lsingle PUnit.unit := diff --git a/Mathlib/RingTheory/AdjoinRoot.lean b/Mathlib/RingTheory/AdjoinRoot.lean index bba7f1e3e7457..eed173fefa7a8 100644 --- a/Mathlib/RingTheory/AdjoinRoot.lean +++ b/Mathlib/RingTheory/AdjoinRoot.lean @@ -205,7 +205,6 @@ theorem aeval_eq (p : R[X]) : aeval (root f) p = mk f p := rw [_root_.map_mul, aeval_C, map_pow, aeval_X, RingHom.map_mul, mk_C, RingHom.map_pow, mk_X] rfl --- Porting note: the following proof was partly in term-mode, but I was not able to fix it. theorem adjoinRoot_eq_top : Algebra.adjoin R ({root f} : Set (AdjoinRoot f)) = ⊤ := by refine Algebra.eq_top_iff.2 fun x => ?_ induction x using AdjoinRoot.induction_on with @@ -402,7 +401,6 @@ def modByMonicHom (hg : g.Monic) : AdjoinRoot g →ₗ[R] R[X] := theorem modByMonicHom_mk (hg : g.Monic) (f : R[X]) : modByMonicHom hg (mk g f) = f %ₘ g := rfl --- Porting note: the following proof was partly in term-mode, but I was not able to fix it. theorem mk_leftInverse (hg : g.Monic) : Function.LeftInverse (mk g) (modByMonicHom hg) := by intro f induction f using AdjoinRoot.induction_on @@ -716,7 +714,6 @@ def quotAdjoinRootEquivQuotPolynomialQuot : ((Ideal.quotEquivOfEq (by rw [map_span, Set.image_singleton])).trans (Polynomial.quotQuotEquivComm I f).symm)) --- Porting note: mathlib3 proof was a long `rw` that timeouts. @[simp] theorem quotAdjoinRootEquivQuotPolynomialQuot_mk_of (p : R[X]) : quotAdjoinRootEquivQuotPolynomialQuot I f (Ideal.Quotient.mk (I.map (of f)) (mk f p)) = diff --git a/Mathlib/RingTheory/WittVector/IsPoly.lean b/Mathlib/RingTheory/WittVector/IsPoly.lean index 69617badd6694..87bd44d9849c4 100644 --- a/Mathlib/RingTheory/WittVector/IsPoly.lean +++ b/Mathlib/RingTheory/WittVector/IsPoly.lean @@ -311,15 +311,11 @@ end ZeroOne /-- Addition of Witt vectors is a polynomial function. -/ -- Porting note: replaced `@[is_poly]` with `instance`. instance addIsPoly₂ [Fact p.Prime] : IsPoly₂ p fun _ _ => (· + ·) := - -- porting note: the proof was - -- `⟨⟨wittAdd p, by intros; dsimp only [WittVector.hasAdd]; simp [eval]⟩⟩` ⟨⟨wittAdd p, by intros; ext; exact add_coeff _ _ _⟩⟩ /-- Multiplication of Witt vectors is a polynomial function. -/ -- Porting note: replaced `@[is_poly]` with `instance`. instance mulIsPoly₂ [Fact p.Prime] : IsPoly₂ p fun _ _ => (· * ·) := - -- porting note: the proof was - -- `⟨⟨wittMul p, by intros; dsimp only [WittVector.hasMul]; simp [eval]⟩⟩` ⟨⟨wittMul p, by intros; ext; exact mul_coeff _ _ _⟩⟩ -- unfortunately this is not universe polymorphic, merely because `f` isn't diff --git a/Mathlib/Topology/Algebra/Monoid.lean b/Mathlib/Topology/Algebra/Monoid.lean index 126c6636dd313..a938067727496 100644 --- a/Mathlib/Topology/Algebra/Monoid.lean +++ b/Mathlib/Topology/Algebra/Monoid.lean @@ -302,8 +302,6 @@ theorem isClosed_setOf_map_mul [Mul M₁] [Mul M₂] [ContinuousMul M₂] : isClosed_iInter fun x => isClosed_iInter fun y => isClosed_eq (continuous_apply _) - -- Porting note: proof was: - -- `((continuous_apply _).mul (continuous_apply _))` (by continuity) -- Porting note: split variables command over two lines, can't change explicitness at the same time diff --git a/Mathlib/Topology/Constructions.lean b/Mathlib/Topology/Constructions.lean index 646dbbbcf8c5a..e0bb583947c37 100644 --- a/Mathlib/Topology/Constructions.lean +++ b/Mathlib/Topology/Constructions.lean @@ -880,11 +880,9 @@ theorem continuous_isRight : Continuous (isRight : X ⊕ Y → Bool) := continuous_sum_dom.2 ⟨continuous_const, continuous_const⟩ @[continuity, fun_prop] --- Porting note: the proof was `continuous_sup_rng_left continuous_coinduced_rng` theorem continuous_inl : Continuous (@inl X Y) := ⟨fun _ => And.left⟩ @[continuity, fun_prop] --- Porting note: the proof was `continuous_sup_rng_right continuous_coinduced_rng` theorem continuous_inr : Continuous (@inr X Y) := ⟨fun _ => And.right⟩ @[fun_prop, continuity] diff --git a/Mathlib/Topology/FiberBundle/Constructions.lean b/Mathlib/Topology/FiberBundle/Constructions.lean index a9ac5e6eb5149..5ecf7487e2490 100644 --- a/Mathlib/Topology/FiberBundle/Constructions.lean +++ b/Mathlib/Topology/FiberBundle/Constructions.lean @@ -250,11 +250,8 @@ universe u v w₁ w₂ U variable {B : Type u} (F : Type v) (E : B → Type w₁) {B' : Type w₂} (f : B' → B) -instance [∀ x : B, TopologicalSpace (E x)] : ∀ x : B', TopologicalSpace ((f *ᵖ E) x) := by - -- Porting note: Original proof was `delta_instance Bundle.Pullback` - intro x - rw [Bundle.Pullback] - infer_instance +instance [∀ x : B, TopologicalSpace (E x)] : ∀ x : B', TopologicalSpace ((f *ᵖ E) x) := + inferInstanceAs (∀ x, TopologicalSpace (E (f x))) variable [TopologicalSpace B'] [TopologicalSpace (TotalSpace F E)] From a4d283b5f2f1bfb9f5c6d8c8c441d5e9d3b3373b Mon Sep 17 00:00:00 2001 From: Peter Nelson <71660771+apnelson1@users.noreply.github.com> Date: Thu, 17 Oct 2024 19:56:29 +0000 Subject: [PATCH 112/505] feat(Order/BooleanAlgebra, Data/Set/Basic): diff_le_diff_iff (#17876) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR adds a lemma stating that, for `x, y ≤ z` in a boolean algebra, we have `z \ x ≤ z \ y` iff `y ≤ x`. We also add the same lemma for sets/finsets. --- Mathlib/Data/Finset/Basic.lean | 4 ++++ Mathlib/Data/Set/Basic.lean | 4 ++++ Mathlib/Order/BooleanAlgebra.lean | 5 +++++ 3 files changed, 13 insertions(+) diff --git a/Mathlib/Data/Finset/Basic.lean b/Mathlib/Data/Finset/Basic.lean index d8f077efcc2ea..f226586bced60 100644 --- a/Mathlib/Data/Finset/Basic.lean +++ b/Mathlib/Data/Finset/Basic.lean @@ -1828,6 +1828,10 @@ theorem sdiff_empty : s \ ∅ = s := theorem sdiff_subset_sdiff (hst : s ⊆ t) (hvu : v ⊆ u) : s \ u ⊆ t \ v := sdiff_le_sdiff hst hvu +theorem sdiff_subset_sdiff_iff_subset {r : Finset α} (hs : s ⊆ r) (ht : t ⊆ r) : + r \ s ⊆ r \ t ↔ t ⊆ s := + sdiff_le_sdiff_iff_le hs ht + @[simp, norm_cast] theorem coe_sdiff (s₁ s₂ : Finset α) : ↑(s₁ \ s₂) = (s₁ \ s₂ : Set α) := Set.ext fun _ => mem_sdiff diff --git a/Mathlib/Data/Set/Basic.lean b/Mathlib/Data/Set/Basic.lean index 2cc7a256bcfdf..63e22e76e74a8 100644 --- a/Mathlib/Data/Set/Basic.lean +++ b/Mathlib/Data/Set/Basic.lean @@ -1497,6 +1497,10 @@ theorem diff_subset_diff_left {s₁ s₂ t : Set α} (h : s₁ ⊆ s₂) : s₁ theorem diff_subset_diff_right {s t u : Set α} (h : t ⊆ u) : s \ u ⊆ s \ t := sdiff_le_sdiff_left ‹t ≤ u› +theorem diff_subset_diff_iff_subset {r : Set α} (hs : s ⊆ r) (ht : t ⊆ r) : + r \ s ⊆ r \ t ↔ t ⊆ s := + sdiff_le_sdiff_iff_le hs ht + theorem compl_eq_univ_diff (s : Set α) : sᶜ = univ \ s := top_sdiff.symm diff --git a/Mathlib/Order/BooleanAlgebra.lean b/Mathlib/Order/BooleanAlgebra.lean index 60726797718c4..ba0cb4558f27b 100644 --- a/Mathlib/Order/BooleanAlgebra.lean +++ b/Mathlib/Order/BooleanAlgebra.lean @@ -359,6 +359,11 @@ theorem sdiff_eq_comm (hy : y ≤ x) (hz : z ≤ x) : x \ y = z ↔ x \ z = y := theorem eq_of_sdiff_eq_sdiff (hxz : x ≤ z) (hyz : y ≤ z) (h : z \ x = z \ y) : x = y := by rw [← sdiff_sdiff_eq_self hxz, h, sdiff_sdiff_eq_self hyz] +theorem sdiff_le_sdiff_iff_le (hx : x ≤ z) (hy : y ≤ z) : z \ x ≤ z \ y ↔ y ≤ x := by + refine ⟨fun h ↦ ?_, sdiff_le_sdiff_left⟩ + rw [← sdiff_sdiff_eq_self hx, ← sdiff_sdiff_eq_self hy] + exact sdiff_le_sdiff_left h + theorem sdiff_sdiff_left' : (x \ y) \ z = x \ y ⊓ x \ z := by rw [sdiff_sdiff_left, sdiff_sup] theorem sdiff_sdiff_sup_sdiff : z \ (x \ y ⊔ y \ x) = z ⊓ (z \ x ⊔ y) ⊓ (z \ y ⊔ x) := From 78d236b37033084670ab8e8b7b0514f719bc6696 Mon Sep 17 00:00:00 2001 From: Antoine Chambert-Loir Date: Thu, 17 Oct 2024 20:34:59 +0000 Subject: [PATCH 113/505] feat(GroupTheory/GroupAction/Blocks): lemmas to handle blocks on finite sets (#14029) This PR proves specific properties of blocks when the action is on a finite set Co-authored-by: Thomas Browning --- Mathlib/GroupTheory/GroupAction/Blocks.lean | 138 ++++++++++++++++++-- Mathlib/GroupTheory/Index.lean | 20 ++- 2 files changed, 147 insertions(+), 11 deletions(-) diff --git a/Mathlib/GroupTheory/GroupAction/Blocks.lean b/Mathlib/GroupTheory/GroupAction/Blocks.lean index dddd660c93991..e88c5972e23aa 100644 --- a/Mathlib/GroupTheory/GroupAction/Blocks.lean +++ b/Mathlib/GroupTheory/GroupAction/Blocks.lean @@ -4,11 +4,14 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Antoine Chambert-Loir -/ +import Mathlib.Algebra.Group.Subgroup.Actions +import Mathlib.Data.Set.Card import Mathlib.Data.Setoid.Partition import Mathlib.GroupTheory.GroupAction.Basic import Mathlib.GroupTheory.GroupAction.Pointwise import Mathlib.GroupTheory.GroupAction.SubMulAction -import Mathlib.Algebra.Group.Subgroup.Actions +import Mathlib.GroupTheory.Index +import Mathlib.Tactic.IntervalCases /-! # Blocks @@ -22,9 +25,21 @@ Given `SMul G X`, an action of a type `G` on a type `X`, we define The non-existence of nontrivial blocks is the definition of primitive actions. +## Results for actions on finite sets + +- `IsBlock.ncard_block_mul_ncard_orbit_eq` : The cardinality of a block +multiplied by the number of its translates is the cardinal of the ambient type + +- `IsBlock.is_univ_of_large_block` : a too large block is equal to `Set.univ` + +- `IsBlock.is_subsingleton` : a too small block is a subsingleton + +- `IsBlock.of_subset` : the intersections of the translates of a finite subset +that contain a given point is a block + ## References -We follow [wieland1964]. +We follow [wielandt1964]. -/ @@ -200,12 +215,13 @@ theorem isBlock_orbit (a : X) : IsBlock G (orbit G a) := variable (X) /-- The full set is a (trivial) block -/ -theorem isFixedBlock_top : IsFixedBlock G (⊤ : Set X) := - fun _ ↦ by simp only [Set.top_eq_univ, Set.smul_set_univ] +theorem isFixedBlock_univ : IsFixedBlock G (Set.univ : Set X) := + fun _ ↦ by simp only [Set.smul_set_univ] +@[deprecated (since := "2024-09-14")] alias isFixedBlock_top := isFixedBlock_univ /-- The full set is a (trivial) block -/ -theorem isBlock_top : IsBlock G (⊤ : Set X) := - (isFixedBlock_top _).isBlock +theorem isBlock_univ : IsBlock G (Set.univ : Set X) := + (isFixedBlock_univ _).isBlock variable {X} @@ -282,9 +298,9 @@ theorem IsBlock.inter {B₁ B₂ : Set X} (h₁ : IsBlock G B₁) (h₂ : IsBloc theorem IsBlock.iInter {ι : Type*} {B : ι → Set X} (hB : ∀ i : ι, IsBlock G (B i)) : IsBlock G (⋂ i, B i) := by by_cases hι : (IsEmpty ι) - · -- ι = ∅, block = ⊤ - suffices (⋂ i : ι, B i) = Set.univ by simpa only [this] using isBlock_top X - simpa only [Set.top_eq_univ, Set.iInter_eq_univ] using (hι.elim' ·) + · -- ι = ∅, block = univ + suffices (⋂ i : ι, B i) = Set.univ by simpa only [this] using isBlock_univ X + simpa only [Set.iInter_eq_univ] using (hι.elim' ·) rw [IsBlock.def_one] intro g rw [Set.smul_set_iInter] @@ -470,6 +486,110 @@ def block_stabilizerOrderIso [htGX : IsPretransitive G X] (a : X) : end Stabilizer +section Finite + +namespace IsBlock + +variable [IsPretransitive G X] {B : Set X} + +theorem ncard_block_eq_relindex (hB : IsBlock G B) {x : X} (hx : x ∈ B) : + B.ncard = (stabilizer G x).relindex (stabilizer G B) := by + have key : (stabilizer G x).subgroupOf (stabilizer G B) = stabilizer (stabilizer G B) x := by + ext; rfl + rw [Subgroup.relindex, key, index_stabilizer, hB.orbit_stabilizer_eq hx] + +/-- The cardinality of the ambient space is the product of the cardinality of a block + by the cardinality of the set of translates of that block -/ +theorem ncard_block_mul_ncard_orbit_eq (hB : IsBlock G B) (hB_ne : B.Nonempty) : + Set.ncard B * Set.ncard (orbit G B) = Nat.card X := by + obtain ⟨x, hx⟩ := hB_ne + rw [ncard_block_eq_relindex hB hx, ← index_stabilizer, + Subgroup.relindex_mul_index (hB.stabilizer_le hx), index_stabilizer_of_transitive] + +/-- The cardinality of a block divides the cardinality of the ambient type -/ +theorem ncard_dvd_card (hB : IsBlock G B) (hB_ne : B.Nonempty) : + Set.ncard B ∣ Nat.card X := + Dvd.intro _ (hB.ncard_block_mul_ncard_orbit_eq hB_ne) + +/-- A too large block is equal to `univ` -/ +theorem eq_univ_card_lt [hX : Finite X] (hB : IsBlock G B) (hB' : Nat.card X < Set.ncard B * 2) : + B = Set.univ := by + rcases Set.eq_empty_or_nonempty B with rfl | hB_ne + · simp only [Set.ncard_empty, zero_mul, not_lt_zero'] at hB' + have key := hB.ncard_block_mul_ncard_orbit_eq hB_ne + rw [← key, mul_lt_mul_iff_of_pos_left (by rwa [Set.ncard_pos])] at hB' + interval_cases (orbit G B).ncard + · rw [mul_zero, eq_comm, Nat.card_eq_zero, or_iff_left hX.not_infinite] at key + exact (IsEmpty.exists_iff.mp hB_ne).elim + · rw [mul_one, ← Set.ncard_univ] at key + rw [Set.eq_of_subset_of_ncard_le (Set.subset_univ B) key.ge] + +/-- If a block has too many translates, then it is a (sub)singleton -/ +theorem subsingleton_of_card_lt [Finite X] (hB : IsBlock G B) + (hB' : Nat.card X < 2 * Set.ncard (orbit G B)) : + B.Subsingleton := by + suffices Set.ncard B < 2 by + rw [Nat.lt_succ_iff, Set.ncard_le_one_iff_eq] at this + cases this with + | inl h => rw [h]; exact Set.subsingleton_empty + | inr h => + obtain ⟨a, ha⟩ := h; rw [ha]; exact Set.subsingleton_singleton + cases Set.eq_empty_or_nonempty B with + | inl h => rw [h, Set.ncard_empty]; norm_num + | inr h => + rw [← hB.ncard_block_mul_ncard_orbit_eq h, lt_iff_not_ge] at hB' + rw [← not_le] + exact fun hb ↦ hB' (Nat.mul_le_mul_right _ hb) + +/- The assumption `B.Finite` is necessary : + For G = ℤ acting on itself, a = 0 and B = ℕ, the translates `k • B` of the statement + are just `k + ℕ`, for `k ≤ 0`, and the corresponding intersection is `ℕ`, which is not a block. + (Remark by Thomas Browning) -/ +-- Note : add {B} because otherwise Lean includes `hB : IsBlock G B` +/-- The intersection of the translates of a *finite* subset which contain a given point +is a block (Wielandt, th. 7.3 )-/ +theorem of_subset {B : Set X} (a : X) (hfB : B.Finite) : + IsBlock G (⋂ (k : G) (_ : a ∈ k • B), k • B) := by + let B' := ⋂ (k : G) (_ : a ∈ k • B), k • B + cases' Set.eq_empty_or_nonempty B with hfB_e hfB_ne + · simp [hfB_e, isBlock_univ] + have hB'₀ : ∀ (k : G) (_ : a ∈ k • B), B' ≤ k • B := by + intro k hk + exact Set.biInter_subset_of_mem hk + have hfB' : B'.Finite := by + obtain ⟨b, hb : b ∈ B⟩ := hfB_ne + obtain ⟨k, hk : k • b = a⟩ := exists_smul_eq G b a + apply Set.Finite.subset (Set.Finite.map _ hfB) (hB'₀ k ⟨b, hb, hk⟩) + have hag : ∀ g : G, a ∈ g • B' → B' ≤ g • B' := by + intro g hg x hx + -- a = g • b; b ∈ B'; a ∈ k • B → b ∈ k • B + simp only [B', Set.mem_iInter, Set.mem_smul_set_iff_inv_smul_mem, + smul_smul, ← mul_inv_rev] at hg hx ⊢ + exact fun _ ↦ hx _ ∘ hg _ + have hag' (g : G) (hg : a ∈ g • B') : B' = g • B' := by + rw [eq_comm, ← mem_stabilizer_iff, mem_stabilizer_of_finite_iff_le_smul _ hfB'] + exact hag g hg + rw [mk_notempty_one] + intro g hg + rw [← Set.nonempty_iff_ne_empty] at hg + obtain ⟨b : X, hb' : b ∈ g • B', hb : b ∈ B'⟩ := Set.nonempty_def.mp hg + obtain ⟨k : G, hk : k • a = b⟩ := exists_smul_eq G a b + have hak : a ∈ k⁻¹ • B' := by + refine ⟨b, hb, ?_⟩ + simp only [← hk, inv_smul_smul] + have hagk : a ∈ (k⁻¹ * g) • B' := by + rw [mul_smul, Set.mem_smul_set_iff_inv_smul_mem, inv_inv, hk] + exact hb' + have hkB' : B' = k⁻¹ • B' := hag' k⁻¹ hak + have hgkB' : B' = (k⁻¹ * g) • B' := hag' (k⁻¹ * g) hagk + rw [mul_smul] at hgkB' + rw [← smul_eq_iff_eq_inv_smul] at hkB' hgkB' + rw [← hgkB', hkB'] + +end IsBlock + +end Finite + end Group end MulAction diff --git a/Mathlib/GroupTheory/Index.lean b/Mathlib/GroupTheory/Index.lean index 99284c35f6952..977fdac89d177 100644 --- a/Mathlib/GroupTheory/Index.lean +++ b/Mathlib/GroupTheory/Index.lean @@ -3,8 +3,9 @@ Copyright (c) 2021 Thomas Browning. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Thomas Browning -/ -import Mathlib.Data.Finite.Card import Mathlib.Algebra.BigOperators.GroupWithZero.Finset +import Mathlib.Data.Finite.Card +import Mathlib.Data.Set.Card import Mathlib.GroupTheory.Coset.Card import Mathlib.GroupTheory.Finiteness import Mathlib.GroupTheory.GroupAction.Quotient @@ -30,7 +31,7 @@ Several theorems proved in this file are known as Lagrange's theorem. - `relindex_mul_index` : If `H ≤ K`, then `H.relindex K * K.index = H.index` - `index_dvd_of_le` : If `H ≤ K`, then `K.index ∣ H.index` - `relindex_mul_relindex` : `relindex` is multiplicative in towers - +- `MulAction.index_stabilizer`: the index of the stabilizer is the cardinality of the orbit -/ @@ -558,6 +559,21 @@ end FiniteIndex end Subgroup +namespace MulAction + +variable (G : Type*) {X : Type*} [Group G] [MulAction G X] (x : X) + +theorem index_stabilizer : + (stabilizer G x).index = (orbit G x).ncard := + (Nat.card_congr (MulAction.orbitEquivQuotientStabilizer G x)).symm.trans + (Set.Nat.card_coe_set_eq (orbit G x)) + +theorem index_stabilizer_of_transitive [IsPretransitive G X] : + (stabilizer G x).index = Nat.card X := by + rw [index_stabilizer, orbit_eq_univ, Set.ncard_univ] + +end MulAction + namespace MonoidHom open Finset From 34de7267074d0f2532903e3e72770b04bfa15ae5 Mon Sep 17 00:00:00 2001 From: jouglasheen Date: Thu, 17 Oct 2024 20:35:01 +0000 Subject: [PATCH 114/505] feat: a nonarchimedean group is totally disconnected (#16687) Prove that a nonarchimedean group is a totally disconnected topological space. Co-authored-by: Kevin Buzzard Co-authored-by: jouglasheen <159258007+jouglasheen@users.noreply.github.com> --- Mathlib.lean | 1 + .../Nonarchimedean/TotallyDisconnected.lean | 60 +++++++++++++++++++ 2 files changed, 61 insertions(+) create mode 100644 Mathlib/Topology/Algebra/Nonarchimedean/TotallyDisconnected.lean diff --git a/Mathlib.lean b/Mathlib.lean index 91da140ade63c..d7f45c2607328 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4555,6 +4555,7 @@ import Mathlib.Topology.Algebra.Nonarchimedean.AdicTopology import Mathlib.Topology.Algebra.Nonarchimedean.Bases import Mathlib.Topology.Algebra.Nonarchimedean.Basic import Mathlib.Topology.Algebra.Nonarchimedean.Completion +import Mathlib.Topology.Algebra.Nonarchimedean.TotallyDisconnected import Mathlib.Topology.Algebra.OpenSubgroup import Mathlib.Topology.Algebra.Order.Archimedean import Mathlib.Topology.Algebra.Order.Compact diff --git a/Mathlib/Topology/Algebra/Nonarchimedean/TotallyDisconnected.lean b/Mathlib/Topology/Algebra/Nonarchimedean/TotallyDisconnected.lean new file mode 100644 index 0000000000000..fd36afa2ef8c6 --- /dev/null +++ b/Mathlib/Topology/Algebra/Nonarchimedean/TotallyDisconnected.lean @@ -0,0 +1,60 @@ +/- +Copyright (c) 2024 Jou Glasheen. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jou Glasheen, Kevin Buzzard, David Loeffler, Yongle Hu, Johan Commelin +-/ + +import Mathlib.Topology.Algebra.Nonarchimedean.Basic + +/-! +# Total separatedness of nonarchimedean groups + +In this file, we prove that a nonarchimedean group is a totally separated topological space. +The fact that a nonarchimedean group is a totally disconnected topological space +is implied by the fact that a nonarchimedean group is totally separated. + +## Main results + +- `NonarchimedeanGroup.instTotallySeparated`: + A nonarchimedean group is a totally separated topological space. + +## Notation + + - `G` : Is a nonarchimedean group. + - `V` : Is an open subgroup which is a neighbourhood of the identity in `G`. + +## References + +See Proposition 2.3.9 and Problem 63 in [F. Q. Gouvêa, *p-adic numbers*][gouvea1997]. +-/ + +open Pointwise TopologicalSpace + +variable {G : Type*} [TopologicalSpace G] [Group G] [NonarchimedeanGroup G] [T2Space G] + +namespace NonarchimedeanGroup + +@[to_additive] +lemma exists_openSubgroup_separating {a b : G} (h : a ≠ b) : + ∃ (V : OpenSubgroup G), Disjoint (a • (V : Set G)) (b • V) := by + obtain ⟨u, v, _, open_v, mem_u, mem_v, dis⟩ := t2_separation (h ∘ inv_mul_eq_one.mp) + obtain ⟨V, hV⟩ := is_nonarchimedean v (open_v.mem_nhds mem_v) + use V + simp only [Disjoint, Set.le_eq_subset, Set.bot_eq_empty, Set.subset_empty_iff] + intros x mem_aV mem_bV + by_contra! con + obtain ⟨s, hs⟩ := con + have hsa : s ∈ a • (V : Set G) := mem_aV hs + have hsb : s ∈ b • (V : Set G) := mem_bV hs + rw [mem_leftCoset_iff] at hsa hsb + refine dis.subset_compl_right mem_u (hV ?_) + simpa [mul_assoc] using mul_mem hsa (inv_mem hsb) + +@[to_additive] +instance (priority := 100) instTotallySeparated : TotallySeparatedSpace G where + isTotallySeparated_univ x _ y _ hxy := by + obtain ⟨V, dxy⟩ := exists_openSubgroup_separating hxy + exact ⟨_, _, V.isOpen.smul x, (V.isClosed.smul x).isOpen_compl, mem_own_leftCoset .., + dxy.subset_compl_left <| mem_own_leftCoset .., by simp, disjoint_compl_right⟩ + +end NonarchimedeanGroup From 389e09795accaebbea904dccedfed6322feb5b86 Mon Sep 17 00:00:00 2001 From: Matthew Robert Ballard Date: Thu, 17 Oct 2024 20:59:58 +0000 Subject: [PATCH 115/505] chore(Algebra.Order.Monoid.TypeTags): split into unbundled and bundled ordered algebra (#16047) We move the unbundled ordered algebra results out of `Algebra.Order.Monoid.TypeTags` into `Algebra.Order.Monoid.Unbundled.TypeTags` to avoid importing bundled ordered algebra results unless necessary. --- Mathlib.lean | 1 + Mathlib/Algebra/Order/Monoid/ToMulBot.lean | 2 +- Mathlib/Algebra/Order/Monoid/TypeTags.lean | 104 +--------------- .../Order/Monoid/Unbundled/TypeTags.lean | 111 ++++++++++++++++++ Mathlib/Algebra/Order/SuccPred/TypeTags.lean | 2 +- 5 files changed, 116 insertions(+), 104 deletions(-) create mode 100644 Mathlib/Algebra/Order/Monoid/Unbundled/TypeTags.lean diff --git a/Mathlib.lean b/Mathlib.lean index d7f45c2607328..721ac852ea637 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -653,6 +653,7 @@ import Mathlib.Algebra.Order.Monoid.Unbundled.ExistsOfLE import Mathlib.Algebra.Order.Monoid.Unbundled.MinMax import Mathlib.Algebra.Order.Monoid.Unbundled.OrderDual import Mathlib.Algebra.Order.Monoid.Unbundled.Pow +import Mathlib.Algebra.Order.Monoid.Unbundled.TypeTags import Mathlib.Algebra.Order.Monoid.Unbundled.WithTop import Mathlib.Algebra.Order.Monoid.Units import Mathlib.Algebra.Order.Monoid.WithTop diff --git a/Mathlib/Algebra/Order/Monoid/ToMulBot.lean b/Mathlib/Algebra/Order/Monoid/ToMulBot.lean index 96f8542fa024c..ab578ee5f849e 100644 --- a/Mathlib/Algebra/Order/Monoid/ToMulBot.lean +++ b/Mathlib/Algebra/Order/Monoid/ToMulBot.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro, Johannes Hölzl -/ import Mathlib.Algebra.Order.GroupWithZero.Canonical -import Mathlib.Algebra.Order.Monoid.TypeTags +import Mathlib.Algebra.Order.Monoid.Unbundled.TypeTags import Mathlib.Algebra.Group.Equiv.Basic import Mathlib.Algebra.Order.Monoid.Unbundled.WithTop diff --git a/Mathlib/Algebra/Order/Monoid/TypeTags.lean b/Mathlib/Algebra/Order/Monoid/TypeTags.lean index c8a194531eb5b..31b5a417097b5 100644 --- a/Mathlib/Algebra/Order/Monoid/TypeTags.lean +++ b/Mathlib/Algebra/Order/Monoid/TypeTags.lean @@ -3,62 +3,13 @@ Copyright (c) 2016 Jeremy Avigad. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro, Johannes Hölzl -/ -import Mathlib.Algebra.Group.TypeTags -import Mathlib.Algebra.Order.Monoid.Defs +import Mathlib.Algebra.Order.Monoid.Unbundled.TypeTags import Mathlib.Algebra.Order.Monoid.Canonical.Defs -/-! # Ordered monoid structures on `Multiplicative α` and `Additive α`. -/ +/-! # Bundled ordered monoid structures on `Multiplicative α` and `Additive α`. -/ variable {α : Type*} -instance : ∀ [LE α], LE (Multiplicative α) := - fun {inst} => inst - -instance : ∀ [LE α], LE (Additive α) := - fun {inst} => inst - -instance : ∀ [LT α], LT (Multiplicative α) := - fun {inst} => inst - -instance : ∀ [LT α], LT (Additive α) := - fun {inst} => inst - -instance Multiplicative.preorder : ∀ [Preorder α], Preorder (Multiplicative α) := - fun {inst} => inst - -instance Additive.preorder : ∀ [Preorder α], Preorder (Additive α) := - fun {inst} => inst - -instance Multiplicative.partialOrder : ∀ [PartialOrder α], PartialOrder (Multiplicative α) := - fun {inst} => inst - -instance Additive.partialOrder : ∀ [PartialOrder α], PartialOrder (Additive α) := - fun {inst} => inst - -instance Multiplicative.linearOrder : ∀ [LinearOrder α], LinearOrder (Multiplicative α) := - fun {inst} => inst - -instance Additive.linearOrder : ∀ [LinearOrder α], LinearOrder (Additive α) := - fun {inst} => inst - -instance Multiplicative.orderBot [LE α] : ∀ [OrderBot α], OrderBot (Multiplicative α) := - fun {inst} => inst - -instance Additive.orderBot [LE α] : ∀ [OrderBot α], OrderBot (Additive α) := - fun {inst} => inst - -instance Multiplicative.orderTop [LE α] : ∀ [OrderTop α], OrderTop (Multiplicative α) := - fun {inst} => inst - -instance Additive.orderTop [LE α] : ∀ [OrderTop α], OrderTop (Additive α) := - fun {inst} => inst - -instance Multiplicative.boundedOrder [LE α] : ∀ [BoundedOrder α], BoundedOrder (Multiplicative α) := - fun {inst} => inst - -instance Additive.boundedOrder [LE α] : ∀ [BoundedOrder α], BoundedOrder (Additive α) := - fun {inst} => inst - instance Multiplicative.orderedCommMonoid [OrderedAddCommMonoid α] : OrderedCommMonoid (Multiplicative α) := { Multiplicative.partialOrder, Multiplicative.commMonoid with @@ -87,13 +38,6 @@ instance Additive.linearOrderedAddCommMonoid [LinearOrderedCommMonoid α] : LinearOrderedAddCommMonoid (Additive α) := { Additive.linearOrder, Additive.orderedAddCommMonoid with } -instance Multiplicative.existsMulOfLe [Add α] [LE α] [ExistsAddOfLE α] : - ExistsMulOfLE (Multiplicative α) := - ⟨@exists_add_of_le α _ _ _⟩ - -instance Additive.existsAddOfLe [Mul α] [LE α] [ExistsMulOfLE α] : ExistsAddOfLE (Additive α) := - ⟨@exists_mul_of_le α _ _ _⟩ - instance Multiplicative.canonicallyOrderedCommMonoid [CanonicallyOrderedAddCommMonoid α] : CanonicallyOrderedCommMonoid (Multiplicative α) := { Multiplicative.orderedCommMonoid, Multiplicative.orderBot, @@ -112,47 +56,3 @@ instance Multiplicative.canonicallyLinearOrderedCommMonoid instance [CanonicallyLinearOrderedCommMonoid α] : CanonicallyLinearOrderedAddCommMonoid (Additive α) := { Additive.canonicallyOrderedAddCommMonoid, Additive.linearOrder with } - -namespace Additive - -variable [Preorder α] - -@[simp] -theorem ofMul_le {a b : α} : ofMul a ≤ ofMul b ↔ a ≤ b := - Iff.rfl - -@[simp] -theorem ofMul_lt {a b : α} : ofMul a < ofMul b ↔ a < b := - Iff.rfl - -@[simp] -theorem toMul_le {a b : Additive α} : toMul a ≤ toMul b ↔ a ≤ b := - Iff.rfl - -@[simp] -theorem toMul_lt {a b : Additive α} : toMul a < toMul b ↔ a < b := - Iff.rfl - -end Additive - -namespace Multiplicative - -variable [Preorder α] - -@[simp] -theorem ofAdd_le {a b : α} : ofAdd a ≤ ofAdd b ↔ a ≤ b := - Iff.rfl - -@[simp] -theorem ofAdd_lt {a b : α} : ofAdd a < ofAdd b ↔ a < b := - Iff.rfl - -@[simp] -theorem toAdd_le {a b : Multiplicative α} : toAdd a ≤ toAdd b ↔ a ≤ b := - Iff.rfl - -@[simp] -theorem toAdd_lt {a b : Multiplicative α} : toAdd a < toAdd b ↔ a < b := - Iff.rfl - -end Multiplicative diff --git a/Mathlib/Algebra/Order/Monoid/Unbundled/TypeTags.lean b/Mathlib/Algebra/Order/Monoid/Unbundled/TypeTags.lean new file mode 100644 index 0000000000000..1647d43c6a434 --- /dev/null +++ b/Mathlib/Algebra/Order/Monoid/Unbundled/TypeTags.lean @@ -0,0 +1,111 @@ +/- +Copyright (c) 2016 Jeremy Avigad. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro, Johannes Hölzl +-/ +import Mathlib.Algebra.Group.TypeTags +import Mathlib.Algebra.Order.Monoid.Unbundled.ExistsOfLE +import Mathlib.Order.BoundedOrder + +/-! # Ordered monoid structures on `Multiplicative α` and `Additive α`. -/ + +variable {α : Type*} + +instance : ∀ [LE α], LE (Multiplicative α) := + fun {inst} => inst + +instance : ∀ [LE α], LE (Additive α) := + fun {inst} => inst + +instance : ∀ [LT α], LT (Multiplicative α) := + fun {inst} => inst + +instance : ∀ [LT α], LT (Additive α) := + fun {inst} => inst + +instance Multiplicative.preorder : ∀ [Preorder α], Preorder (Multiplicative α) := + fun {inst} => inst + +instance Additive.preorder : ∀ [Preorder α], Preorder (Additive α) := + fun {inst} => inst + +instance Multiplicative.partialOrder : ∀ [PartialOrder α], PartialOrder (Multiplicative α) := + fun {inst} => inst + +instance Additive.partialOrder : ∀ [PartialOrder α], PartialOrder (Additive α) := + fun {inst} => inst + +instance Multiplicative.linearOrder : ∀ [LinearOrder α], LinearOrder (Multiplicative α) := + fun {inst} => inst + +instance Additive.linearOrder : ∀ [LinearOrder α], LinearOrder (Additive α) := + fun {inst} => inst + +instance Multiplicative.orderBot [LE α] : ∀ [OrderBot α], OrderBot (Multiplicative α) := + fun {inst} => inst + +instance Additive.orderBot [LE α] : ∀ [OrderBot α], OrderBot (Additive α) := + fun {inst} => inst + +instance Multiplicative.orderTop [LE α] : ∀ [OrderTop α], OrderTop (Multiplicative α) := + fun {inst} => inst + +instance Additive.orderTop [LE α] : ∀ [OrderTop α], OrderTop (Additive α) := + fun {inst} => inst + +instance Multiplicative.boundedOrder [LE α] : ∀ [BoundedOrder α], BoundedOrder (Multiplicative α) := + fun {inst} => inst + +instance Additive.boundedOrder [LE α] : ∀ [BoundedOrder α], BoundedOrder (Additive α) := + fun {inst} => inst + +instance Multiplicative.existsMulOfLe [Add α] [LE α] [ExistsAddOfLE α] : + ExistsMulOfLE (Multiplicative α) := + ⟨@exists_add_of_le α _ _ _⟩ + +instance Additive.existsAddOfLe [Mul α] [LE α] [ExistsMulOfLE α] : ExistsAddOfLE (Additive α) := + ⟨@exists_mul_of_le α _ _ _⟩ + +namespace Additive + +variable [Preorder α] + +@[simp] +theorem ofMul_le {a b : α} : ofMul a ≤ ofMul b ↔ a ≤ b := + Iff.rfl + +@[simp] +theorem ofMul_lt {a b : α} : ofMul a < ofMul b ↔ a < b := + Iff.rfl + +@[simp] +theorem toMul_le {a b : Additive α} : toMul a ≤ toMul b ↔ a ≤ b := + Iff.rfl + +@[simp] +theorem toMul_lt {a b : Additive α} : toMul a < toMul b ↔ a < b := + Iff.rfl + +end Additive + +namespace Multiplicative + +variable [Preorder α] + +@[simp] +theorem ofAdd_le {a b : α} : ofAdd a ≤ ofAdd b ↔ a ≤ b := + Iff.rfl + +@[simp] +theorem ofAdd_lt {a b : α} : ofAdd a < ofAdd b ↔ a < b := + Iff.rfl + +@[simp] +theorem toAdd_le {a b : Multiplicative α} : toAdd a ≤ toAdd b ↔ a ≤ b := + Iff.rfl + +@[simp] +theorem toAdd_lt {a b : Multiplicative α} : toAdd a < toAdd b ↔ a < b := + Iff.rfl + +end Multiplicative diff --git a/Mathlib/Algebra/Order/SuccPred/TypeTags.lean b/Mathlib/Algebra/Order/SuccPred/TypeTags.lean index 4293cd851694a..7c69e6dc40162 100644 --- a/Mathlib/Algebra/Order/SuccPred/TypeTags.lean +++ b/Mathlib/Algebra/Order/SuccPred/TypeTags.lean @@ -3,8 +3,8 @@ Copyright (c) 2024 Yakov Pechersky. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yakov Pechersky -/ -import Mathlib.Algebra.Order.Monoid.TypeTags import Mathlib.Order.SuccPred.Archimedean +import Mathlib.Algebra.Order.Monoid.Unbundled.TypeTags /-! # Successor and predecessor on type tags From 23a1a4bd2117e0e4a6391122bcf08d7c2edf6008 Mon Sep 17 00:00:00 2001 From: Christian Merten Date: Thu, 17 Oct 2024 22:32:03 +0000 Subject: [PATCH 116/505] feat(AlgebraicGeometry): proper morphisms of schemes (#17863) We define proper morphisms of schemes and show standard stability properties. Partly from the valuative criterion project. Co-authored by: Andrew Yang --- Mathlib.lean | 1 + .../Morphisms/ClosedImmersion.lean | 14 ++++ .../Morphisms/FiniteType.lean | 3 + .../AlgebraicGeometry/Morphisms/Proper.lean | 69 +++++++++++++++++++ .../Morphisms/UniversallyClosed.lean | 12 +++- 5 files changed, 98 insertions(+), 1 deletion(-) create mode 100644 Mathlib/AlgebraicGeometry/Morphisms/Proper.lean diff --git a/Mathlib.lean b/Mathlib.lean index 721ac852ea637..6871945d645ea 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -868,6 +868,7 @@ import Mathlib.AlgebraicGeometry.Morphisms.FiniteType import Mathlib.AlgebraicGeometry.Morphisms.IsIso import Mathlib.AlgebraicGeometry.Morphisms.OpenImmersion import Mathlib.AlgebraicGeometry.Morphisms.Preimmersion +import Mathlib.AlgebraicGeometry.Morphisms.Proper import Mathlib.AlgebraicGeometry.Morphisms.QuasiCompact import Mathlib.AlgebraicGeometry.Morphisms.QuasiSeparated import Mathlib.AlgebraicGeometry.Morphisms.RingHomProperties diff --git a/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean b/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean index c78e3b86d608d..932325058e426 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean @@ -6,6 +6,7 @@ Authors: Amelia Livingston, Christian Merten, Jonas van der Schaaf import Mathlib.AlgebraicGeometry.Morphisms.Preimmersion import Mathlib.AlgebraicGeometry.Morphisms.QuasiSeparated import Mathlib.AlgebraicGeometry.Morphisms.RingHomProperties +import Mathlib.AlgebraicGeometry.Morphisms.FiniteType import Mathlib.AlgebraicGeometry.ResidueField import Mathlib.RingTheory.RingHom.Surjective @@ -269,4 +270,17 @@ lemma IsClosedImmersion.stableUnderBaseChange : exact ⟨inferInstance, RingHom.surjective_stableUnderBaseChange.pullback_fst_app_top _ RingHom.surjective_respectsIso f _ hsurj⟩ +/-- Closed immersions are locally of finite type. -/ +instance (priority := 900) {X Y : Scheme.{u}} (f : X ⟶ Y) [h : IsClosedImmersion f] : + LocallyOfFiniteType f := by + wlog hY : IsAffine Y + · rw [IsLocalAtTarget.iff_of_iSup_eq_top (P := @LocallyOfFiniteType) _ + (iSup_affineOpens_eq_top Y)] + intro U + have H : IsClosedImmersion (f ∣_ U) := IsLocalAtTarget.restrict h U + exact this _ U.2 + obtain ⟨_, hf⟩ := h.isAffine_surjective_of_isAffine + rw [HasRingHomProperty.iff_of_isAffine (P := @LocallyOfFiniteType)] + exact RingHom.FiniteType.of_surjective (Scheme.Hom.app f ⊤) hf + end AlgebraicGeometry diff --git a/Mathlib/AlgebraicGeometry/Morphisms/FiniteType.lean b/Mathlib/AlgebraicGeometry/Morphisms/FiniteType.lean index def22435c69c0..a5bbcc0abb5f2 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/FiniteType.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/FiniteType.lean @@ -59,6 +59,9 @@ theorem locallyOfFiniteType_of_comp {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) [LocallyOfFiniteType (f ≫ g)] : LocallyOfFiniteType f := HasRingHomProperty.of_comp (fun _ _ ↦ RingHom.FiniteType.of_comp_finiteType) ‹_› +instance : MorphismProperty.IsMultiplicative @LocallyOfFiniteType where + id_mem _ := inferInstance + open scoped TensorProduct in lemma locallyOfFiniteType_stableUnderBaseChange : MorphismProperty.StableUnderBaseChange @LocallyOfFiniteType := diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Proper.lean b/Mathlib/AlgebraicGeometry/Morphisms/Proper.lean new file mode 100644 index 0000000000000..b54565bfb3af8 --- /dev/null +++ b/Mathlib/AlgebraicGeometry/Morphisms/Proper.lean @@ -0,0 +1,69 @@ +/- +Copyright (c) 2024 Christian Merten, Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Christian Merten, Andrew Yang +-/ +import Mathlib.AlgebraicGeometry.Morphisms.Separated +import Mathlib.AlgebraicGeometry.Morphisms.UniversallyClosed +import Mathlib.AlgebraicGeometry.Morphisms.FiniteType + +/-! + +# Proper morphisms + +A morphism of schemes is proper if it is separated, universally closed and (locally) of finite type. +Note that we don't require quasi-compact, since this is implied by universally closed (TODO). + +-/ + + +noncomputable section + +open CategoryTheory + +universe u + +namespace AlgebraicGeometry + +variable {X Y : Scheme.{u}} (f : X ⟶ Y) + +/-- A morphism is proper if it is separated, universally closed and locally of finite type. -/ +@[mk_iff] +class IsProper extends IsSeparated f, UniversallyClosed f, LocallyOfFiniteType f : Prop where + +lemma isProper_eq : @IsProper = + (@IsSeparated ⊓ @UniversallyClosed : MorphismProperty Scheme) ⊓ @LocallyOfFiniteType := by + ext X Y f + rw [isProper_iff, ← and_assoc] + rfl + +namespace IsProper + +instance : MorphismProperty.RespectsIso @IsProper := by + rw [isProper_eq] + infer_instance + +instance stableUnderComposition : MorphismProperty.IsStableUnderComposition @IsProper := by + rw [isProper_eq] + infer_instance + +instance : MorphismProperty.IsMultiplicative @IsProper := by + rw [isProper_eq] + infer_instance + +instance (priority := 900) [IsClosedImmersion f] : IsProper f where + +lemma stableUnderBaseChange : MorphismProperty.StableUnderBaseChange @IsProper := by + rw [isProper_eq] + exact MorphismProperty.StableUnderBaseChange.inf + (MorphismProperty.StableUnderBaseChange.inf + IsSeparated.stableUnderBaseChange universallyClosed_stableUnderBaseChange) + locallyOfFiniteType_stableUnderBaseChange + +instance : IsLocalAtTarget @IsProper := by + rw [isProper_eq] + infer_instance + +end IsProper + +end AlgebraicGeometry diff --git a/Mathlib/AlgebraicGeometry/Morphisms/UniversallyClosed.lean b/Mathlib/AlgebraicGeometry/Morphisms/UniversallyClosed.lean index 6a5489d992476..166651638bcaa 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/UniversallyClosed.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/UniversallyClosed.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Andrew Yang. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang -/ -import Mathlib.AlgebraicGeometry.Morphisms.Constructors +import Mathlib.AlgebraicGeometry.Morphisms.ClosedImmersion import Mathlib.Topology.LocalAtTarget /-! @@ -40,6 +40,13 @@ class UniversallyClosed (f : X ⟶ Y) : Prop where theorem universallyClosed_eq : @UniversallyClosed = universally (topologically @IsClosedMap) := by ext X Y f; rw [universallyClosed_iff] +instance (priority := 900) [IsClosedImmersion f] : UniversallyClosed f := by + rw [universallyClosed_eq] + intro X' Y' i₁ i₂ f' hf + have hf' : IsClosedImmersion f' := + IsClosedImmersion.stableUnderBaseChange hf.flip inferInstance + exact hf'.base_closed.isClosedMap + theorem universallyClosed_respectsIso : RespectsIso @UniversallyClosed := universallyClosed_eq.symm ▸ universally_respectsIso (topologically @IsClosedMap) @@ -59,6 +66,9 @@ instance universallyClosedTypeComp {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) [hf : UniversallyClosed f] [hg : UniversallyClosed g] : UniversallyClosed (f ≫ g) := comp_mem _ _ _ hf hg +instance : MorphismProperty.IsMultiplicative @UniversallyClosed where + id_mem _ := inferInstance + instance universallyClosed_fst {X Y Z : Scheme} (f : X ⟶ Z) (g : Y ⟶ Z) [hg : UniversallyClosed g] : UniversallyClosed (pullback.fst f g) := universallyClosed_stableUnderBaseChange.fst f g hg From f1ebf24bc990d6d810d4780353eb74d9ecb39803 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Fri, 18 Oct 2024 00:49:50 +0000 Subject: [PATCH 117/505] refactor: add `compactSpace_spectrum` and `spectrum_nonempty` fields to the `ContinuousFunctionalCalculus` (#17801) --- Mathlib/Algebra/Algebra/Quasispectrum.lean | 3 +++ .../Instances.lean | 1 + .../NonUnital.lean | 21 +++++++++++++------ .../ContinuousFunctionalCalculus/Order.lean | 12 +++++------ .../Restrict.lean | 10 +++++++++ .../ContinuousFunctionalCalculus/Unital.lean | 7 +++++++ .../ContinuousFunctionalCalculus/Rpow.lean | 8 ++++--- .../Matrix/HermitianFunctionalCalculus.lean | 5 +++++ 8 files changed, 51 insertions(+), 16 deletions(-) diff --git a/Mathlib/Algebra/Algebra/Quasispectrum.lean b/Mathlib/Algebra/Algebra/Quasispectrum.lean index 33a5807e877b8..e94372dc43fdc 100644 --- a/Mathlib/Algebra/Algebra/Quasispectrum.lean +++ b/Mathlib/Algebra/Algebra/Quasispectrum.lean @@ -254,6 +254,9 @@ lemma quasispectrum.not_isUnit_mem (a : A) {r : R} (hr : ¬ IsUnit r) : r ∈ qu lemma quasispectrum.zero_mem [Nontrivial R] (a : A) : 0 ∈ quasispectrum R a := quasispectrum.not_isUnit_mem a <| by simp +theorem quasispectrum.nonempty [Nontrivial R] (a : A) : (quasispectrum R a).Nonempty := + Set.nonempty_of_mem <| quasispectrum.zero_mem R a + instance quasispectrum.instZero [Nontrivial R] (a : A) : Zero (quasispectrum R a) where zero := ⟨0, quasispectrum.zero_mem R a⟩ diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean index 50abcebd079d7..ea67a1f84c970 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean @@ -155,6 +155,7 @@ instance IsStarNormal.instContinuousFunctionalCalculus {A : Type*} [NormedRing A [CStarRing A] [CompleteSpace A] [NormedAlgebra ℂ A] [StarModule ℂ A] : ContinuousFunctionalCalculus ℂ (IsStarNormal : A → Prop) where predicate_zero := isStarNormal_zero + spectrum_nonempty a _ := spectrum.nonempty a exists_cfc_of_predicate a ha := by refine ⟨(elementalStarAlgebra ℂ a).subtype.comp <| continuousFunctionalCalculus a, ?hom_closedEmbedding, ?hom_id, ?hom_map_spectrum, ?predicate_hom⟩ diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean index 361cdc8aefebc..d4a2d892c661e 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean @@ -73,11 +73,16 @@ class NonUnitalContinuousFunctionalCalculus (R : Type*) {A : Type*} (p : outPara [ContinuousStar R] [NonUnitalRing A] [StarRing A] [TopologicalSpace A] [Module R A] [IsScalarTower R A A] [SMulCommClass R A A] : Prop where predicate_zero : p 0 + [compactSpace_quasispectrum : ∀ a : A, CompactSpace (σₙ R a)] exists_cfc_of_predicate : ∀ a, p a → ∃ φ : C(σₙ R a, R)₀ →⋆ₙₐ[R] A, ClosedEmbedding φ ∧ φ ⟨(ContinuousMap.id R).restrict <| σₙ R a, rfl⟩ = a ∧ (∀ f, σₙ R (φ f) = Set.range f) ∧ ∀ f, p (φ f) --- TODO: try to unify with the unital version. The `ℝ≥0` case makes it tricky. +-- this instance should not be activated everywhere but it is useful when developing generic API +-- for the continuous functional calculus +scoped[NonUnitalContinuousFunctionalCalculus] +attribute [instance] NonUnitalContinuousFunctionalCalculus.compactSpace_quasispectrum + /-- A class guaranteeing that the non-unital continuous functional calculus is uniquely determined by the properties that it is a continuous non-unital star algebra homomorphism mapping the (restriction of) the identity to `a`. This is the necessary tool used to establish `cfcₙHom_comp` @@ -685,7 +690,10 @@ lemma cfcₙHom_of_cfcHom_map_quasispectrum {a : A} (ha : p a) : simp only [Set.mem_singleton_iff] at hx ⊢ rw [show x = 0 from Subtype.val_injective hx, map_zero] -variable [CompleteSpace R] [h_cpct : ∀ a : A, CompactSpace (spectrum R a)] +variable [CompleteSpace R] + +-- gives access to the `ContinuousFunctionalCalculus.compactSpace_spectrum` instance +open scoped ContinuousFunctionalCalculus lemma closedEmbedding_cfcₙHom_of_cfcHom {a : A} (ha : p a) : ClosedEmbedding (cfcₙHom_of_cfcHom R ha) := by @@ -709,6 +717,10 @@ lemma closedEmbedding_cfcₙHom_of_cfcHom {a : A} (ha : p a) : instance ContinuousFunctionalCalculus.toNonUnital : NonUnitalContinuousFunctionalCalculus R p where predicate_zero := cfc_predicate_zero R + compactSpace_quasispectrum a := by + have h_cpct : CompactSpace (spectrum R a) := inferInstance + simp only [← isCompact_iff_compactSpace, quasispectrum_eq_spectrum_union_zero] at h_cpct ⊢ + exact h_cpct |>.union isCompact_singleton exists_cfc_of_predicate _ ha := ⟨cfcₙHom_of_cfcHom R ha, closedEmbedding_cfcₙHom_of_cfcHom ha, @@ -716,12 +728,9 @@ instance ContinuousFunctionalCalculus.toNonUnital : NonUnitalContinuousFunctiona cfcₙHom_of_cfcHom_map_quasispectrum ha, fun _ ↦ cfcHom_predicate ha _⟩ +open scoped NonUnitalContinuousFunctionalCalculus in lemma cfcₙHom_eq_cfcₙHom_of_cfcHom [UniqueNonUnitalContinuousFunctionalCalculus R A] {a : A} (ha : p a) : cfcₙHom (R := R) ha = cfcₙHom_of_cfcHom R ha := by - have h_cpct' : CompactSpace (σₙ R a) := by - specialize h_cpct a - simp_rw [← isCompact_iff_compactSpace, quasispectrum_eq_spectrum_union_zero] at h_cpct ⊢ - exact h_cpct.union isCompact_singleton refine UniqueNonUnitalContinuousFunctionalCalculus.eq_of_continuous_of_map_id (σₙ R a) ?_ _ _ ?_ ?_ ?_ · simp diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Order.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Order.lean index 8536353d9475b..10e0978f35a5c 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Order.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Order.lean @@ -85,20 +85,20 @@ lemma cfc_nnreal_le_iff {A : Type*} [TopologicalSpace A] [Ring A] [StarRing A] [ rw [cfc_nnreal_eq_real, cfc_nnreal_eq_real, cfc_le_iff ..] simp [NNReal.coe_le_coe, ← ha_spec.image] +open ContinuousFunctionalCalculus in /-- In a unital `ℝ`-algebra `A` with a continuous functional calculus, an element `a : A` is larger than some `algebraMap ℝ A r` if and only if every element of the `ℝ`-spectrum is nonnegative. -/ lemma CFC.exists_pos_algebraMap_le_iff {A : Type*} [TopologicalSpace A] [Ring A] [StarRing A] - [PartialOrder A] [StarOrderedRing A] [Algebra ℝ A] [NonnegSpectrumClass ℝ A] + [PartialOrder A] [StarOrderedRing A] [Algebra ℝ A] [NonnegSpectrumClass ℝ A] [Nontrivial A] [ContinuousFunctionalCalculus ℝ (IsSelfAdjoint : A → Prop)] - {a : A} [CompactSpace (spectrum ℝ a)] - (h_non : (spectrum ℝ a).Nonempty) (ha : IsSelfAdjoint a := by cfc_tac) : + {a : A} (ha : IsSelfAdjoint a := by cfc_tac) : (∃ r > 0, algebraMap ℝ A r ≤ a) ↔ (∀ x ∈ spectrum ℝ a, 0 < x) := by have h_cpct : IsCompact (spectrum ℝ a) := isCompact_iff_compactSpace.mpr inferInstance simp_rw [algebraMap_le_iff_le_spectrum (a := a)] refine ⟨?_, fun h ↦ ?_⟩ · rintro ⟨r, hr, hr_le⟩ exact (hr.trans_le <| hr_le · ·) - · obtain ⟨r, hr, hr_min⟩ := h_cpct.exists_isMinOn h_non continuousOn_id + · obtain ⟨r, hr, hr_min⟩ := h_cpct.exists_isMinOn (spectrum_nonempty ℝ a ha) continuousOn_id exact ⟨r, h _ hr, hr_min⟩ section CStar_unital @@ -218,10 +218,8 @@ lemma CStarAlgebra.isUnit_of_le {a b : A} (h₀ : IsUnit a) (ha : 0 ≤ a := by rw [← spectrum.zero_not_mem_iff ℝ≥0] at h₀ ⊢ nontriviality A have hb := (show 0 ≤ a from ha).trans hab - have ha' := IsSelfAdjoint.of_nonneg ha |>.spectrum_nonempty - have hb' := IsSelfAdjoint.of_nonneg hb |>.spectrum_nonempty rw [zero_not_mem_iff, SpectrumRestricts.nnreal_lt_iff (.nnreal_of_nonneg ‹_›), - NNReal.coe_zero, ← CFC.exists_pos_algebraMap_le_iff ‹_›] at h₀ ⊢ + NNReal.coe_zero, ← CFC.exists_pos_algebraMap_le_iff] at h₀ ⊢ peel h₀ with r hr _ exact this.trans hab diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Restrict.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Restrict.lean index cf36028a7922e..bb43228df3fdf 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Restrict.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Restrict.lean @@ -98,6 +98,12 @@ protected theorem cfc (f : C(S, R)) (halg : IsUniformEmbedding (algebraMap R S)) (h : ∀ a, p a ↔ q a ∧ SpectrumRestricts a f) : ContinuousFunctionalCalculus R p where predicate_zero := h0 + spectrum_nonempty a ha := ((h a).mp ha).2.image ▸ + (ContinuousFunctionalCalculus.spectrum_nonempty a ((h a).mp ha).1 |>.image f) + compactSpace_spectrum a := by + have := ContinuousFunctionalCalculus.compactSpace_spectrum (R := S) a + rw [← isCompact_iff_compactSpace] at this ⊢ + simpa using halg.toClosedEmbedding.isCompact_preimage this exists_cfc_of_predicate a ha := by refine ⟨((h a).mp ha).2.starAlgHom (cfcHom ((h a).mp ha).1 (R := S)), ?hom_closedEmbedding, ?hom_id, ?hom_map_spectrum, ?predicate_hom⟩ @@ -235,6 +241,10 @@ protected theorem cfc (f : C(S, R)) (halg : IsUniformEmbedding (algebraMap R S)) (h : ∀ a, p a ↔ q a ∧ QuasispectrumRestricts a f) : NonUnitalContinuousFunctionalCalculus R p where predicate_zero := h0 + compactSpace_quasispectrum a := by + have := NonUnitalContinuousFunctionalCalculus.compactSpace_quasispectrum (R := S) a + rw [← isCompact_iff_compactSpace] at this ⊢ + simpa using halg.toClosedEmbedding.isCompact_preimage this exists_cfc_of_predicate a ha := by refine ⟨((h a).mp ha).2.nonUnitalStarAlgHom (cfcₙHom ((h a).mp ha).1 (R := S)), ?hom_closedEmbedding, ?hom_id, ?hom_map_spectrum, ?predicate_hom⟩ diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean index 8b6473cba0de3..29fa35441c946 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean @@ -162,10 +162,17 @@ class ContinuousFunctionalCalculus (R : Type*) {A : Type*} (p : outParam (A → [CommSemiring R] [StarRing R] [MetricSpace R] [TopologicalSemiring R] [ContinuousStar R] [Ring A] [StarRing A] [TopologicalSpace A] [Algebra R A] : Prop where predicate_zero : p 0 + [compactSpace_spectrum (a : A) : CompactSpace (spectrum R a)] + spectrum_nonempty [Nontrivial A] (a : A) (ha : p a) : (spectrum R a).Nonempty exists_cfc_of_predicate : ∀ a, p a → ∃ φ : C(spectrum R a, R) →⋆ₐ[R] A, ClosedEmbedding φ ∧ φ ((ContinuousMap.id R).restrict <| spectrum R a) = a ∧ (∀ f, spectrum R (φ f) = Set.range f) ∧ ∀ f, p (φ f) +-- this instance should not be activated everywhere but it is useful when developing generic API +-- for the continuous functional calculus +scoped[ContinuousFunctionalCalculus] +attribute [instance] ContinuousFunctionalCalculus.compactSpace_spectrum + /-- A class guaranteeing that the continuous functional calculus is uniquely determined by the properties that it is a continuous star algebra homomorphism mapping the (restriction of) the identity to `a`. This is the necessary tool used to establish `cfcHom_comp` and the more common diff --git a/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/Rpow.lean b/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/Rpow.lean index 9fcf8e77c60a6..fda3885c78a00 100644 --- a/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/Rpow.lean +++ b/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/Rpow.lean @@ -5,8 +5,8 @@ Authors: Frédéric Dupuis -/ import Mathlib.Analysis.SpecialFunctions.Pow.Real -import Mathlib.Analysis.Normed.Algebra.Spectrum import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.NonUnital +import Mathlib.Analysis.SpecialFunctions.Pow.Continuity /-! # Real powers defined via the continuous functional calculus @@ -306,8 +306,10 @@ lemma rpow_intCast (a : Aˣ) (n : ℤ) (ha : (0 : A) ≤ a := by cfc_tac) : section unital_vs_nonunital -variable [∀ (a : A), CompactSpace (spectrum ℝ a)] - [UniqueNonUnitalContinuousFunctionalCalculus ℝ≥0 A] +variable [UniqueNonUnitalContinuousFunctionalCalculus ℝ≥0 A] + +-- provides instance `ContinuousFunctionalCalculus.compactSpace_spectrum` +open scoped ContinuousFunctionalCalculus lemma nnrpow_eq_rpow {a : A} {x : ℝ≥0} (hx : 0 < x) : a ^ x = a ^ (x : ℝ) := by rw [nnrpow_def (A := A), rpow_def, cfcₙ_eq_cfc] diff --git a/Mathlib/LinearAlgebra/Matrix/HermitianFunctionalCalculus.lean b/Mathlib/LinearAlgebra/Matrix/HermitianFunctionalCalculus.lean index ff884d2e289ad..5c9c8e5757541 100644 --- a/Mathlib/LinearAlgebra/Matrix/HermitianFunctionalCalculus.lean +++ b/Mathlib/LinearAlgebra/Matrix/HermitianFunctionalCalculus.lean @@ -135,6 +135,11 @@ instance instContinuousFunctionalCalculus : rw [star_eq_conjTranspose, diagonal_conjTranspose] congr! simp [Pi.star_def, Function.comp_def] + spectrum_nonempty a ha := by + obtain (h | h) := isEmpty_or_nonempty n + · obtain ⟨x, y, hxy⟩ := exists_pair_ne (Matrix n n 𝕜) + exact False.elim <| Matrix.of.symm.injective.ne hxy <| Subsingleton.elim _ _ + · exact eigenvalues_eq_spectrum_real ha ▸ Set.range_nonempty _ predicate_zero := .zero _ instance instUniqueContinuousFunctionalCalculus : From 3f9cdcd1db12e4b662c93305015ce28d909918d8 Mon Sep 17 00:00:00 2001 From: Moritz Firsching Date: Fri, 18 Oct 2024 01:51:14 +0000 Subject: [PATCH 118/505] feat(Tactic/Linter/Lint): improve `dupNamespace` linter (#17631) Refactor the the way the ids are extracted for the `dupNamespace` linter. We also extract this to a new file in anticipation of re-using that function in another linter. The new version finds auto-generated names, such as `to_additive` names, but ignores "internal" names. Like the previous linter, it also supports `export` and `alias`. Co-authored-by: adomani Co-authored-by: damiano Co-authored-by: Moritz Firsching --- Mathlib.lean | 1 + Mathlib/Tactic.lean | 1 + Mathlib/Tactic/DeclarationNames.lean | 45 ++++++++++++++++++++++++++++ Mathlib/Tactic/Linter/Lint.lean | 32 ++++++++------------ test/Lint.lean | 28 +++++++++++++++-- 5 files changed, 85 insertions(+), 22 deletions(-) create mode 100644 Mathlib/Tactic/DeclarationNames.lean diff --git a/Mathlib.lean b/Mathlib.lean index 6871945d645ea..8963ee37e6ec3 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4298,6 +4298,7 @@ import Mathlib.Tactic.Contrapose import Mathlib.Tactic.Conv import Mathlib.Tactic.Convert import Mathlib.Tactic.Core +import Mathlib.Tactic.DeclarationNames import Mathlib.Tactic.DefEqTransformations import Mathlib.Tactic.DeprecateMe import Mathlib.Tactic.DeriveFintype diff --git a/Mathlib/Tactic.lean b/Mathlib/Tactic.lean index c8d7b7ae27494..cd494193a67b4 100644 --- a/Mathlib/Tactic.lean +++ b/Mathlib/Tactic.lean @@ -62,6 +62,7 @@ import Mathlib.Tactic.Contrapose import Mathlib.Tactic.Conv import Mathlib.Tactic.Convert import Mathlib.Tactic.Core +import Mathlib.Tactic.DeclarationNames import Mathlib.Tactic.DefEqTransformations import Mathlib.Tactic.DeprecateMe import Mathlib.Tactic.DeriveFintype diff --git a/Mathlib/Tactic/DeclarationNames.lean b/Mathlib/Tactic/DeclarationNames.lean new file mode 100644 index 0000000000000..3003bab4db21d --- /dev/null +++ b/Mathlib/Tactic/DeclarationNames.lean @@ -0,0 +1,45 @@ +/- +Copyright (c) 2024 Moritz Firsching. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Damiano Testa, Moritz Firsching +-/ + +import Lean.DeclarationRange +import Lean.ResolveName + +/-! +This file contains functions that are used by multiple linters. +-/ + +open Lean Parser Elab Command Meta +namespace Mathlib.Linter + +/-- +If `pos` is a `String.Pos`, then `getNamesFrom pos` returns the array of identifiers +for the names of the declarations whose syntax begins in position at least `pos`. +-/ +def getNamesFrom {m} [Monad m] [MonadEnv m] [MonadFileMap m] (pos : String.Pos) : + m (Array Syntax) := do + let drs := declRangeExt.getState (← getEnv) + let fm ← getFileMap + let mut nms := #[] + for (nm, rgs) in drs do + if pos ≤ fm.ofPosition rgs.range.pos then + let ofPos1 := fm.ofPosition rgs.selectionRange.pos + let ofPos2 := fm.ofPosition rgs.selectionRange.endPos + nms := nms.push (mkIdentFrom (.ofRange ⟨ofPos1, ofPos2⟩) nm) + return nms + +/-- +If `stx` is a syntax node for an `export` statement, then `getAliasSyntax stx` returns the array of +identifiers with the "exported" names. +-/ +def getAliasSyntax {m} [Monad m] [MonadResolveName m] (stx : Syntax) : m (Array Syntax) := do + let mut aliases := #[] + if let `(export $_ ($ids*)) := stx then + let currNamespace ← getCurrNamespace + for idStx in ids do + let id := idStx.getId + aliases := aliases.push + (mkIdentFrom (.ofRange (idStx.raw.getRange?.getD default)) (currNamespace ++ id)) + return aliases diff --git a/Mathlib/Tactic/Linter/Lint.lean b/Mathlib/Tactic/Linter/Lint.lean index 4c0a8fb5df9a4..8641d3a70bb7d 100644 --- a/Mathlib/Tactic/Linter/Lint.lean +++ b/Mathlib/Tactic/Linter/Lint.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn -/ import Batteries.Tactic.Lint +import Mathlib.Tactic.DeclarationNames /-! # Linters for Mathlib @@ -77,29 +78,20 @@ namespace DupNamespaceLinter open Lean Parser Elab Command Meta -/-- `getIds stx` extracts the `declId` nodes from the `Syntax` `stx`. -If `stx` is an `alias` or an `export`, then it extracts an `ident`, instead of a `declId`. -/ -partial -def getIds : Syntax → Array Syntax - | .node _ `Batteries.Tactic.Alias.alias args => args[2:3] - | .node _ ``Lean.Parser.Command.export args => (args[3:4] : Array Syntax).map (·[0]) - | stx@(.node _ _ args) => - ((args.attach.map fun ⟨a, _⟩ ↦ getIds a).foldl (· ++ ·) #[stx]).filter (·.getKind == ``declId) - | _ => default - @[inherit_doc linter.dupNamespace] def dupNamespace : Linter where run := withSetOptionIn fun stx ↦ do if Linter.getLinterValue linter.dupNamespace (← getOptions) then - match getIds stx with - | #[id] => - let ns := (← getScope).currNamespace - let declName := ns ++ (if id.getKind == ``declId then id[0].getId else id.getId) - let nm := declName.components - let some (dup, _) := nm.zip (nm.tailD []) |>.find? fun (x, y) ↦ x == y - | return - Linter.logLint linter.dupNamespace id - m!"The namespace '{dup}' is duplicated in the declaration '{declName}'" - | _ => return + let mut aliases := #[] + if let some exp := stx.find? (·.isOfKind `Lean.Parser.Command.export) then + aliases ← getAliasSyntax exp + for id in (← getNamesFrom (stx.getPos?.getD default)) ++ aliases do + let declName := id.getId + if declName.hasMacroScopes then continue + let nm := declName.components + let some (dup, _) := nm.zip (nm.tailD []) |>.find? fun (x, y) ↦ x == y + | continue + Linter.logLint linter.dupNamespace id + m!"The namespace '{dup}' is duplicated in the declaration '{declName}'" initialize addLinter dupNamespace diff --git a/test/Lint.lean b/test/Lint.lean index 25bf9c9147d13..7f6b2c917ddfb 100644 --- a/test/Lint.lean +++ b/test/Lint.lean @@ -7,6 +7,23 @@ import Mathlib.Order.SetNotation set_option linter.dupNamespace false +namespace termG +-- this creates a hygienic declaration starting with `termG.termG.«_@».test.Lint...` +-- and the linter ignores it +set_option linter.dupNamespace true in +local notation "G" => Unit + +/-- info: [termG, termG] -/ +#guard_msgs in +open Lean in +run_meta + let env ← getEnv + let consts := env.constants.toList.find? (·.1.getRoot == `termG) + let reps := (consts.map (·.1.components.take 2)).getD default + logInfo m!"{reps}" + guard (reps[0]! == reps[1]!) +end termG + /-- warning: The namespace 'add' is duplicated in the declaration 'add.add' note: this linter can be disabled with `set_option linter.dupNamespace false` @@ -25,8 +42,12 @@ note: this linter can be disabled with `set_option linter.dupNamespace false` set_option linter.dupNamespace true in def Foo.foo := True --- the `dupNamespace` linter does not notice that `to_additive` created `Foo.add.add`. +/-- +warning: The namespace 'add' is duplicated in the declaration 'Foo.add.add' +note: this linter can be disabled with `set_option linter.dupNamespace false` +-/ #guard_msgs in +set_option linter.dupNamespace true in @[to_additive] theorem add.mul : True := .intro -- However, the declaration `Foo.add.add` is present in the environment. @@ -52,10 +73,13 @@ namespace add /-- warning: The namespace 'add' is duplicated in the declaration 'add.add' note: this linter can be disabled with `set_option linter.dupNamespace false` +--- +warning: The namespace 'add' is duplicated in the declaration 'add.add' +note: this linter can be disabled with `set_option linter.dupNamespace false` -/ #guard_msgs in set_option linter.dupNamespace true in -export Nat (add) +export Nat (add add_comm add) end add From 1814dc4dcc7a88545743b89f61260b1db6d9fd02 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 18 Oct 2024 02:46:57 +0000 Subject: [PATCH 119/505] refactor(Equiv/Set): use `Disjoint` (#17880) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit ... instead of `s ∩ t ⊆ ∅` in `Equiv.Set.union`. Also add `simps` to `Equiv.Set.univ`. --- Mathlib/Data/Finset/Basic.lean | 2 +- Mathlib/Data/Set/Card.lean | 3 +-- Mathlib/GroupTheory/Perm/Finite.lean | 8 +++---- Mathlib/Logic/Equiv/Set.lean | 32 ++++++++++++--------------- Mathlib/SetTheory/Cardinal/Basic.lean | 2 +- 5 files changed, 21 insertions(+), 26 deletions(-) diff --git a/Mathlib/Data/Finset/Basic.lean b/Mathlib/Data/Finset/Basic.lean index f226586bced60..a94497b850efa 100644 --- a/Mathlib/Data/Finset/Basic.lean +++ b/Mathlib/Data/Finset/Basic.lean @@ -3019,7 +3019,7 @@ open Finset /-- The disjoint union of finsets is a sum -/ def Finset.union (s t : Finset α) (h : Disjoint s t) : s ⊕ t ≃ (s ∪ t : Finset α) := - Equiv.Set.ofEq (coe_union _ _) |>.trans (Equiv.Set.union (disjoint_coe.mpr h).le_bot) |>.symm + Equiv.Set.ofEq (coe_union _ _) |>.trans (Equiv.Set.union (disjoint_coe.mpr h)) |>.symm @[simp] theorem Finset.union_symm_inl (h : Disjoint s t) (x : s) : diff --git a/Mathlib/Data/Set/Card.lean b/Mathlib/Data/Set/Card.lean index 57aba3430a761..fa03f476c8309 100644 --- a/Mathlib/Data/Set/Card.lean +++ b/Mathlib/Data/Set/Card.lean @@ -109,8 +109,7 @@ theorem encard_ne_zero : s.encard ≠ 0 ↔ s.Nonempty := by theorem encard_union_eq (h : Disjoint s t) : (s ∪ t).encard = s.encard + t.encard := by classical - have e := (Equiv.Set.union (by rwa [subset_empty_iff, ← disjoint_iff_inter_eq_empty])).symm - simp [encard, ← PartENat.card_congr e, PartENat.card_sum, PartENat.withTopEquiv] + simp [encard, PartENat.card_congr (Equiv.Set.union h), PartENat.card_sum, PartENat.withTopEquiv] theorem encard_insert_of_not_mem {a : α} (has : a ∉ s) : (insert a s).encard = s.encard + 1 := by rw [← union_singleton, encard_union_eq (by simpa), encard_singleton] diff --git a/Mathlib/GroupTheory/Perm/Finite.lean b/Mathlib/GroupTheory/Perm/Finite.lean index e7ead8ca2a46c..8c9cb242cb255 100644 --- a/Mathlib/GroupTheory/Perm/Finite.lean +++ b/Mathlib/GroupTheory/Perm/Finite.lean @@ -184,9 +184,9 @@ theorem Disjoint.isConj_mul [Finite α] {σ τ π ρ : Perm α} (hc1 : IsConj σ have hd2'' := disjoint_coe.2 (disjoint_iff_disjoint_support.1 hd2) refine isConj_of_support_equiv ?_ ?_ · refine - ((Equiv.Set.ofEq hd1').trans (Equiv.Set.union hd1''.le_bot)).trans + ((Equiv.Set.ofEq hd1').trans (Equiv.Set.union hd1'')).trans ((Equiv.sumCongr (subtypeEquiv f fun a => ?_) (subtypeEquiv g fun a => ?_)).trans - ((Equiv.Set.ofEq hd2').trans (Equiv.Set.union hd2''.le_bot)).symm) <;> + ((Equiv.Set.ofEq hd2').trans (Equiv.Set.union hd2'')).symm) <;> · simp only [Set.mem_image, toEmbedding_apply, exists_eq_right, support_conj, coe_map, apply_eq_iff_eq] · intro x hx @@ -195,7 +195,7 @@ theorem Disjoint.isConj_mul [Finite α] {σ τ π ρ : Perm α} (hc1 : IsConj σ rw [hd1', Set.mem_union] at hx cases' hx with hxσ hxτ · rw [mem_coe, mem_support] at hxσ - rw [Set.union_apply_left hd1''.le_bot _, Set.union_apply_left hd1''.le_bot _] + rw [Set.union_apply_left, Set.union_apply_left] · simp only [subtypeEquiv_apply, Perm.coe_mul, Sum.map_inl, comp_apply, Set.union_symm_apply_left, Subtype.coe_mk, apply_eq_iff_eq] have h := (hd2 (f x)).resolve_left ?_ @@ -206,7 +206,7 @@ theorem Disjoint.isConj_mul [Finite α] {σ τ π ρ : Perm α} (hc1 : IsConj σ · rwa [Subtype.coe_mk, Perm.mul_apply, (hd1 x).resolve_left hxσ, mem_coe, apply_mem_support, mem_support] · rw [mem_coe, ← apply_mem_support, mem_support] at hxτ - rw [Set.union_apply_right hd1''.le_bot _, Set.union_apply_right hd1''.le_bot _] + rw [Set.union_apply_right, Set.union_apply_right] · simp only [subtypeEquiv_apply, Perm.coe_mul, Sum.map_inr, comp_apply, Set.union_symm_apply_right, Subtype.coe_mk, apply_eq_iff_eq] have h := (hd2 (g (τ x))).resolve_right ?_ diff --git a/Mathlib/Logic/Equiv/Set.lean b/Mathlib/Logic/Equiv/Set.lean index fb24c904db068..a734cb9993db0 100644 --- a/Mathlib/Logic/Equiv/Set.lean +++ b/Mathlib/Logic/Equiv/Set.lean @@ -174,8 +174,8 @@ def image {α β : Type*} (e : α ≃ β) (s : Set α) : namespace Set --- Porting note: Removed attribute @[simps apply symm_apply] /-- `univ α` is equivalent to `α`. -/ +@[simps apply symm_apply] protected def univ (α) : @univ α ≃ α := ⟨Subtype.val, fun a => ⟨a, trivial⟩, fun ⟨_, _⟩ => rfl, fun _ => rfl⟩ @@ -203,25 +203,25 @@ protected def union' {α} {s t : Set α} (p : α → Prop) [DecidablePred p] (hs rcases o with (⟨x, h⟩ | ⟨x, h⟩) <;> [simp [hs _ h]; simp [ht _ h]] /-- If sets `s` and `t` are disjoint, then `s ∪ t` is equivalent to `s ⊕ t`. -/ -protected def union {α} {s t : Set α} [DecidablePred fun x => x ∈ s] (H : s ∩ t ⊆ ∅) : +protected def union {α} {s t : Set α} [DecidablePred fun x => x ∈ s] (H : Disjoint s t) : (s ∪ t : Set α) ≃ s ⊕ t := - Set.union' (fun x => x ∈ s) (fun _ => id) fun _ xt xs => H ⟨xs, xt⟩ + Set.union' (fun x => x ∈ s) (fun _ => id) fun _ xt xs => Set.disjoint_left.mp H xs xt -theorem union_apply_left {α} {s t : Set α} [DecidablePred fun x => x ∈ s] (H : s ∩ t ⊆ ∅) +theorem union_apply_left {α} {s t : Set α} [DecidablePred fun x => x ∈ s] (H : Disjoint s t) {a : (s ∪ t : Set α)} (ha : ↑a ∈ s) : Equiv.Set.union H a = Sum.inl ⟨a, ha⟩ := dif_pos ha -theorem union_apply_right {α} {s t : Set α} [DecidablePred fun x => x ∈ s] (H : s ∩ t ⊆ ∅) +theorem union_apply_right {α} {s t : Set α} [DecidablePred fun x => x ∈ s] (H : Disjoint s t) {a : (s ∪ t : Set α)} (ha : ↑a ∈ t) : Equiv.Set.union H a = Sum.inr ⟨a, ha⟩ := - dif_neg fun h => H ⟨h, ha⟩ + dif_neg fun h => Set.disjoint_left.mp H h ha @[simp] -theorem union_symm_apply_left {α} {s t : Set α} [DecidablePred fun x => x ∈ s] (H : s ∩ t ⊆ ∅) +theorem union_symm_apply_left {α} {s t : Set α} [DecidablePred fun x => x ∈ s] (H : Disjoint s t) (a : s) : (Equiv.Set.union H).symm (Sum.inl a) = ⟨a, by simp⟩ := rfl @[simp] -theorem union_symm_apply_right {α} {s t : Set α} [DecidablePred fun x => x ∈ s] (H : s ∩ t ⊆ ∅) +theorem union_symm_apply_right {α} {s t : Set α} [DecidablePred fun x => x ∈ s] (H : Disjoint s t) (a : t) : (Equiv.Set.union H).symm (Sum.inr a) = ⟨a, by simp⟩ := rfl @@ -247,7 +247,7 @@ protected def insert {α} {s : Set.{u} α} [DecidablePred (· ∈ s)] {a : α} ( (insert a s : Set α) ≃ s ⊕ PUnit.{u + 1} := calc (insert a s : Set α) ≃ ↥(s ∪ {a}) := Equiv.Set.ofEq (by simp) - _ ≃ s ⊕ ({a} : Set α) := Equiv.Set.union fun x ⟨hx, _⟩ => by simp_all + _ ≃ s ⊕ ({a} : Set α) := Equiv.Set.union <| by simpa _ ≃ s ⊕ PUnit.{u + 1} := sumCongr (Equiv.refl _) (Equiv.Set.singleton _) @[simp] @@ -273,7 +273,7 @@ theorem insert_apply_right {α} {s : Set.{u} α} [DecidablePred (· ∈ s)] {a : /-- If `s : Set α` is a set with decidable membership, then `s ⊕ sᶜ` is equivalent to `α`. -/ protected def sumCompl {α} (s : Set α) [DecidablePred (· ∈ s)] : s ⊕ (sᶜ : Set α) ≃ α := calc - s ⊕ (sᶜ : Set α) ≃ ↥(s ∪ sᶜ) := (Equiv.Set.union (by simp [Set.ext_iff])).symm + s ⊕ (sᶜ : Set α) ≃ ↥(s ∪ sᶜ) := (Equiv.Set.union disjoint_compl_right).symm _ ≃ @univ α := Equiv.Set.ofEq (by simp) _ ≃ α := Equiv.Set.univ _ @@ -289,15 +289,11 @@ theorem sumCompl_apply_inr {α : Type u} (s : Set α) [DecidablePred (· ∈ s)] theorem sumCompl_symm_apply_of_mem {α : Type u} {s : Set α} [DecidablePred (· ∈ s)] {x : α} (hx : x ∈ s) : (Equiv.Set.sumCompl s).symm x = Sum.inl ⟨x, hx⟩ := by - have : ((⟨x, Or.inl hx⟩ : (s ∪ sᶜ : Set α)) : α) ∈ s := hx - rw [Equiv.Set.sumCompl] - simpa using Set.union_apply_left (by simp) this + simp [Equiv.Set.sumCompl, Equiv.Set.univ, union_apply_left, hx] theorem sumCompl_symm_apply_of_not_mem {α : Type u} {s : Set α} [DecidablePred (· ∈ s)] {x : α} (hx : x ∉ s) : (Equiv.Set.sumCompl s).symm x = Sum.inr ⟨x, hx⟩ := by - have : ((⟨x, Or.inr hx⟩ : (s ∪ sᶜ : Set α)) : α) ∈ sᶜ := hx - rw [Equiv.Set.sumCompl] - simpa using Set.union_apply_right (by simp) this + simp [Equiv.Set.sumCompl, Equiv.Set.univ, union_apply_right, hx] @[simp] theorem sumCompl_symm_apply {α : Type*} {s : Set α} [DecidablePred (· ∈ s)] {x : s} : @@ -315,7 +311,7 @@ protected def sumDiffSubset {α} {s t : Set α} (h : s ⊆ t) [DecidablePred (· s ⊕ (t \ s : Set α) ≃ t := calc s ⊕ (t \ s : Set α) ≃ (s ∪ t \ s : Set α) := - (Equiv.Set.union (by simp [inter_diff_self])).symm + (Equiv.Set.union disjoint_sdiff_self_right).symm _ ≃ t := Equiv.Set.ofEq (by simp [union_diff_self, union_eq_self_of_subset_left h]) @[simp] @@ -346,7 +342,7 @@ protected def unionSumInter {α : Type u} (s t : Set α) [DecidablePred (· ∈ (s ∪ t : Set α) ⊕ (s ∩ t : Set α) ≃ (s ∪ t \ s : Set α) ⊕ (s ∩ t : Set α) := by rw [union_diff_self] _ ≃ (s ⊕ (t \ s : Set α)) ⊕ (s ∩ t : Set α) := - sumCongr (Set.union <| subset_empty_iff.2 (inter_diff_self _ _)) (Equiv.refl _) + sumCongr (Set.union disjoint_sdiff_self_right) (Equiv.refl _) _ ≃ s ⊕ ((t \ s : Set α) ⊕ (s ∩ t : Set α)) := sumAssoc _ _ _ _ ≃ s ⊕ (t \ s ∪ s ∩ t : Set α) := sumCongr (Equiv.refl _) diff --git a/Mathlib/SetTheory/Cardinal/Basic.lean b/Mathlib/SetTheory/Cardinal/Basic.lean index 8c37e772ce499..22ec5d425d8c1 100644 --- a/Mathlib/SetTheory/Cardinal/Basic.lean +++ b/Mathlib/SetTheory/Cardinal/Basic.lean @@ -1954,7 +1954,7 @@ theorem mk_union_le {α : Type u} (S T : Set α) : #(S ∪ T : Set α) ≤ #S + theorem mk_union_of_disjoint {α : Type u} {S T : Set α} (H : Disjoint S T) : #(S ∪ T : Set α) = #S + #T := by classical - exact Quot.sound ⟨Equiv.Set.union H.le_bot⟩ + exact Quot.sound ⟨Equiv.Set.union H⟩ theorem mk_insert {α : Type u} {s : Set α} {a : α} (h : a ∉ s) : #(insert a s : Set α) = #s + 1 := by From fb09b71265c60a58ebf2c1d4fd6b9477d4506d14 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Fri, 18 Oct 2024 03:51:31 +0000 Subject: [PATCH 120/505] chore: cleanup of pairwise imports (#17888) --- Mathlib/Data/Fin/Basic.lean | 3 +++ Mathlib/Data/List/OfFn.lean | 6 +++--- 2 files changed, 6 insertions(+), 3 deletions(-) diff --git a/Mathlib/Data/Fin/Basic.lean b/Mathlib/Data/Fin/Basic.lean index 715342aa56963..5dbafff567583 100644 --- a/Mathlib/Data/Fin/Basic.lean +++ b/Mathlib/Data/Fin/Basic.lean @@ -626,6 +626,9 @@ theorem leftInverse_cast (eq : n = m) : LeftInverse (cast eq.symm) (cast eq) := theorem rightInverse_cast (eq : n = m) : RightInverse (cast eq.symm) (cast eq) := fun _ => rfl +theorem cast_lt_cast (eq : n = m) {a b : Fin n} : cast eq a < cast eq b ↔ a < b := + Iff.rfl + theorem cast_le_cast (eq : n = m) {a b : Fin n} : cast eq a ≤ cast eq b ↔ a ≤ b := Iff.rfl diff --git a/Mathlib/Data/List/OfFn.lean b/Mathlib/Data/List/OfFn.lean index a3271be86ffd8..27001ba3ea4b0 100644 --- a/Mathlib/Data/List/OfFn.lean +++ b/Mathlib/Data/List/OfFn.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ import Batteries.Data.List.OfFn -import Batteries.Data.List.Pairwise import Mathlib.Data.Fin.Tuple.Basic /-! @@ -178,8 +177,9 @@ theorem ofFn_fin_repeat {m} (a : Fin m → α) (n : ℕ) : @[simp] theorem pairwise_ofFn {R : α → α → Prop} {n} {f : Fin n → α} : (ofFn f).Pairwise R ↔ ∀ ⦃i j⦄, i < j → R (f i) (f j) := by - simp only [pairwise_iff_get, (Fin.rightInverse_cast (length_ofFn f)).surjective.forall, get_ofFn, - ← Fin.not_le, Fin.cast_le_cast] + simp only [pairwise_iff_getElem, length_ofFn, List.getElem_ofFn, + (Fin.rightInverse_cast (length_ofFn f)).surjective.forall, Fin.forall_iff, Fin.cast_mk, + Fin.mk_lt_mk, forall_comm (α := (_ : Prop)) (β := ℕ)] /-- Lists are equivalent to the sigma type of tuples of a given length. -/ @[simps] From 077e0efd658ccaa95474fb130feeded70c602f6a Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Fri, 18 Oct 2024 04:24:40 +0000 Subject: [PATCH 121/505] chore: move deprecated file Mathlib/Data/ByteArray into Mathlib/Deprecated (#17884) --- Mathlib.lean | 2 +- Mathlib/{Data => Deprecated}/ByteArray.lean | 0 scripts/noshake.json | 3 ++- 3 files changed, 3 insertions(+), 2 deletions(-) rename Mathlib/{Data => Deprecated}/ByteArray.lean (100%) diff --git a/Mathlib.lean b/Mathlib.lean index 8963ee37e6ec3..9a66b6063fc59 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2145,7 +2145,6 @@ import Mathlib.Data.Bool.Count import Mathlib.Data.Bool.Set import Mathlib.Data.Bracket import Mathlib.Data.Bundle -import Mathlib.Data.ByteArray import Mathlib.Data.Char import Mathlib.Data.Complex.Abs import Mathlib.Data.Complex.Basic @@ -2670,6 +2669,7 @@ import Mathlib.Data.ZMod.Quotient import Mathlib.Data.ZMod.Units import Mathlib.Deprecated.AlgebraClasses import Mathlib.Deprecated.Aliases +import Mathlib.Deprecated.ByteArray import Mathlib.Deprecated.Combinator import Mathlib.Deprecated.Equiv import Mathlib.Deprecated.Group diff --git a/Mathlib/Data/ByteArray.lean b/Mathlib/Deprecated/ByteArray.lean similarity index 100% rename from Mathlib/Data/ByteArray.lean rename to Mathlib/Deprecated/ByteArray.lean diff --git a/scripts/noshake.json b/scripts/noshake.json index 681e7f87ee2b5..3e93d378ca0e2 100644 --- a/scripts/noshake.json +++ b/scripts/noshake.json @@ -344,14 +344,15 @@ "Mathlib.Lean.Meta": ["Batteries.Logic"], "Mathlib.Lean.Expr.ExtraRecognizers": ["Mathlib.Data.Set.Operations"], "Mathlib.Lean.Expr.Basic": ["Batteries.Logic"], - "Mathlib.Deprecated.NatLemmas": ["Batteries.Data.Nat.Lemmas", "Batteries.WF"], "Mathlib.GroupTheory.MonoidLocalization.Basic": ["Mathlib.Init.Data.Prod"], "Mathlib.Geometry.Manifold.Sheaf.Smooth": ["Mathlib.CategoryTheory.Sites.Whiskering"], "Mathlib.Geometry.Manifold.PoincareConjecture": ["Mathlib.AlgebraicTopology.FundamentalGroupoid.SimplyConnected", "Mathlib.Util.Superscript"], + "Mathlib.Deprecated.NatLemmas": ["Batteries.Data.Nat.Lemmas", "Batteries.WF"], "Mathlib.Deprecated.MinMax": ["Mathlib.Order.MinMax"], + "Mathlib.Deprecated.ByteArray": ["Batteries.Data.ByteSubarray"], "Mathlib.Data.Vector.Basic": ["Mathlib.Control.Applicative"], "Mathlib.Data.Set.Image": ["Batteries.Tactic.Congr"], "Mathlib.Data.Seq.Parallel": ["Mathlib.Init.Data.Prod"], From 1b9ced20cbaf344811df3cbe11b727d3989bf2e2 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Fri, 18 Oct 2024 04:24:41 +0000 Subject: [PATCH 122/505] chore: cleanup of imports involving List.count (#17885) --- Archive/MiuLanguage/DecisionNec.lean | 2 -- Mathlib/Combinatorics/Enumerative/DyckWord.lean | 1 - Mathlib/Data/List/Count.lean | 3 +-- 3 files changed, 1 insertion(+), 5 deletions(-) diff --git a/Archive/MiuLanguage/DecisionNec.lean b/Archive/MiuLanguage/DecisionNec.lean index 694b599ac899d..3bd06ca5463c6 100644 --- a/Archive/MiuLanguage/DecisionNec.lean +++ b/Archive/MiuLanguage/DecisionNec.lean @@ -5,9 +5,7 @@ Authors: Gihan Marasingha -/ import Archive.MiuLanguage.Basic import Mathlib.Data.List.Basic -import Mathlib.Data.List.Count import Mathlib.Data.Nat.ModEq -import Mathlib.Tactic.Ring /-! # Decision procedure: necessary condition diff --git a/Mathlib/Combinatorics/Enumerative/DyckWord.lean b/Mathlib/Combinatorics/Enumerative/DyckWord.lean index cd58430af7c76..5924d03e54bd1 100644 --- a/Mathlib/Combinatorics/Enumerative/DyckWord.lean +++ b/Mathlib/Combinatorics/Enumerative/DyckWord.lean @@ -3,7 +3,6 @@ Copyright (c) 2024 Jeremy Tan. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Tan -/ -import Batteries.Data.List.Count import Mathlib.Combinatorics.Enumerative.Catalan import Mathlib.Tactic.Positivity diff --git a/Mathlib/Data/List/Count.lean b/Mathlib/Data/List/Count.lean index 89475cec9fe1e..13c127b563b2a 100644 --- a/Mathlib/Data/List/Count.lean +++ b/Mathlib/Data/List/Count.lean @@ -10,8 +10,7 @@ import Mathlib.Tactic.Common # Counting in lists This file proves basic properties of `List.countP` and `List.count`, which count the number of -elements of a list satisfying a predicate and equal to a given element respectively. Their -definitions can be found in `Batteries.Data.List.Basic`. +elements of a list satisfying a predicate and equal to a given element respectively. -/ assert_not_exists Set.range From 4358f93967f3a6a35e2e299ff9de9e7007336549 Mon Sep 17 00:00:00 2001 From: Moritz Firsching Date: Fri, 18 Oct 2024 05:30:33 +0000 Subject: [PATCH 123/505] feat(NumberTheory/FermatNumber): fermatNumber n + 2 (#17659) We provide a simple non-linear recurrence relation for the Fermat numbers. Since this involves subtraction, both Nat and Int versions are given. Co-authored-by: Moritz Firsching --- Mathlib/NumberTheory/Fermat.lean | 30 +++++++++++++++++++++++++++++- 1 file changed, 29 insertions(+), 1 deletion(-) diff --git a/Mathlib/NumberTheory/Fermat.lean b/Mathlib/NumberTheory/Fermat.lean index 5fe9b6b32bc73..848dc49be4810 100644 --- a/Mathlib/NumberTheory/Fermat.lean +++ b/Mathlib/NumberTheory/Fermat.lean @@ -8,6 +8,8 @@ import Mathlib.Algebra.Order.Ring.Basic import Mathlib.Algebra.Order.Star.Basic import Mathlib.Data.Nat.Prime.Defs import Mathlib.Tactic.Ring.RingNF +import Mathlib.Tactic.Linarith + /-! # Fermat numbers @@ -62,6 +64,32 @@ theorem fermatNumber_succ (n : ℕ) : fermatNumber (n + 1) = (fermatNumber n - 1 rw [fermatNumber, pow_succ, mul_comm, Nat.pow_mul'] rfl +theorem two_mul_fermatNumber_sub_one_sq_le_fermatNumber_sq (n : ℕ) : + 2 * (fermatNumber n - 1) ^ 2 ≤ (fermatNumber (n + 1)) ^ 2 := by + simp only [fermatNumber, add_tsub_cancel_right] + have : 0 ≤ 1 + 2 ^ (2 ^ n * 4) := le_add_left 0 (Nat.add 1 _) + ring_nf + linarith + +theorem fermatNumber_eq_fermatNumber_sq_sub_two_mul_fermatNumber_sub_one_sq (n : ℕ) : + fermatNumber (n + 2) = (fermatNumber (n + 1)) ^ 2 - 2 * (fermatNumber n - 1) ^ 2 := by + simp only [fermatNumber, add_sub_self_right] + rw [← add_sub_self_right (2 ^ 2 ^ (n + 2) + 1) <| 2 * 2 ^ 2 ^ (n + 1)] + ring_nf + +end Nat + +open Nat + +theorem Int.fermatNumber_eq_fermatNumber_sq_sub_two_mul_fermatNumber_sub_one_sq (n : ℕ) : + (fermatNumber (n + 2) : ℤ) = (fermatNumber (n + 1)) ^ 2 - 2 * (fermatNumber n - 1) ^ 2 := by + rw [Nat.fermatNumber_eq_fermatNumber_sq_sub_two_mul_fermatNumber_sub_one_sq, + Nat.cast_sub <| two_mul_fermatNumber_sub_one_sq_le_fermatNumber_sq n] + simp only [fermatNumber, push_cast, add_tsub_cancel_right] + +namespace Nat + +open Finset /-- **Goldbach's theorem** : no two distinct Fermat numbers share a common factor greater than one. @@ -78,4 +106,4 @@ theorem coprime_fermatNumber_fermatNumber {k n : ℕ} (h : k ≠ n) : refine ((dvd_prime prime_two).mp h_m).elim id (fun h_two ↦ ?_) exact ((odd_fermatNumber _).not_two_dvd_nat (h_two ▸ h_n)).elim - end Nat +end Nat From 5d0c5274420fa8a2989d403a5058acb316ab06e4 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 18 Oct 2024 05:56:05 +0000 Subject: [PATCH 124/505] feat(PathConnected): review API, add missing lemmas (#17855) --- .../Complex/UpperHalfPlane/Topology.lean | 3 +- Mathlib/Analysis/Convex/Normed.lean | 2 +- Mathlib/Topology/Connected/PathConnected.lean | 204 +++++++++--------- 3 files changed, 103 insertions(+), 106 deletions(-) diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/Topology.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/Topology.lean index 3c32e2e099f23..caecc196aeb8d 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/Topology.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/Topology.lean @@ -55,8 +55,7 @@ instance : T4Space ℍ := inferInstance instance : ContractibleSpace ℍ := (convex_halfspace_im_gt 0).contractibleSpace ⟨I, one_pos.trans_eq I_im.symm⟩ -instance : LocPathConnectedSpace ℍ := - locPathConnected_of_isOpen <| isOpen_lt continuous_const Complex.continuous_im +instance : LocPathConnectedSpace ℍ := openEmbedding_coe.locPathConnectedSpace instance : NoncompactSpace ℍ := by refine ⟨fun h => ?_⟩ diff --git a/Mathlib/Analysis/Convex/Normed.lean b/Mathlib/Analysis/Convex/Normed.lean index 235a27716dee4..fae3e58b55f18 100644 --- a/Mathlib/Analysis/Convex/Normed.lean +++ b/Mathlib/Analysis/Convex/Normed.lean @@ -112,7 +112,7 @@ instance (priority := 100) NormedSpace.instPathConnectedSpace : PathConnectedSpa TopologicalAddGroup.pathConnectedSpace instance (priority := 100) NormedSpace.instLocPathConnectedSpace : LocPathConnectedSpace E := - locPathConnected_of_bases (fun _ => Metric.nhds_basis_ball) fun x r r_pos => + .of_bases (fun _ => Metric.nhds_basis_ball) fun x r r_pos => (convex_ball x r).isPathConnected <| by simp [r_pos] theorem Wbtw.dist_add_dist {x y z : P} (h : Wbtw ℝ x y z) : diff --git a/Mathlib/Topology/Connected/PathConnected.lean b/Mathlib/Topology/Connected/PathConnected.lean index 644fcb694b3a8..e5798309d30d4 100644 --- a/Mathlib/Topology/Connected/PathConnected.lean +++ b/Mathlib/Topology/Connected/PathConnected.lean @@ -785,13 +785,35 @@ theorem Specializes.joinedIn (h : x ⤳ y) (hx : x ∈ F) (hy : y ∈ F) : Joine theorem Inseparable.joinedIn (h : Inseparable x y) (hx : x ∈ F) (hy : y ∈ F) : JoinedIn F x y := h.specializes.joinedIn hx hy -/-! ### Path component -/ +theorem JoinedIn.map_continuousOn (h : JoinedIn F x y) {f : X → Y} (hf : ContinuousOn f F) : + JoinedIn (f '' F) (f x) (f y) := + let ⟨γ, hγ⟩ := h + ⟨γ.map' <| hf.mono (range_subset_iff.mpr hγ), fun t ↦ mem_image_of_mem _ (hγ t)⟩ + +theorem JoinedIn.map (h : JoinedIn F x y) {f : X → Y} (hf : Continuous f) : + JoinedIn (f '' F) (f x) (f y) := + h.map_continuousOn hf.continuousOn + +theorem Inducing.joinedIn_image {f : X → Y} (hf : Inducing f) (hx : x ∈ F) (hy : y ∈ F) : + JoinedIn (f '' F) (f x) (f y) ↔ JoinedIn F x y := by + refine ⟨?_, (.map · hf.continuous)⟩ + rintro ⟨γ, hγ⟩ + choose γ' hγ'F hγ' using hγ + have h₀ : x ⤳ γ' 0 := by rw [← hf.specializes_iff, hγ', γ.source] + have h₁ : γ' 1 ⤳ y := by rw [← hf.specializes_iff, hγ', γ.target] + have h : JoinedIn F (γ' 0) (γ' 1) := by + refine ⟨⟨⟨γ', ?_⟩, rfl, rfl⟩, hγ'F⟩ + simpa only [hf.continuous_iff, comp_def, hγ'] using map_continuous γ + exact (h₀.joinedIn hx (hγ'F _)).trans <| h.trans <| h₁.joinedIn (hγ'F _) hy +/-! ### Path component -/ /-- The path component of `x` is the set of points that can be joined to `x`. -/ def pathComponent (x : X) := { y | Joined x y } +theorem mem_pathComponent_iff : x ∈ pathComponent y ↔ Joined y x := .rfl + @[simp] theorem mem_pathComponent_self (x : X) : x ∈ pathComponent x := Joined.refl x @@ -867,24 +889,16 @@ theorem IsPathConnected.image' (hF : IsPathConnected F) exact hf.mono (range_subset_iff.2 (hx y_in).somePath_mem) /-- If `f` is continuous and `F` is path-connected, so is `f(F)`. -/ -theorem IsPathConnected.image (hF : IsPathConnected F) {f : X → Y} - (hf : Continuous f) : IsPathConnected (f '' F) := hF.image' hf.continuousOn +theorem IsPathConnected.image (hF : IsPathConnected F) {f : X → Y} (hf : Continuous f) : + IsPathConnected (f '' F) := + hF.image' hf.continuousOn /-- If `f : X → Y` is a `Inducing`, `f(F)` is path-connected iff `F` is. -/ nonrec theorem Inducing.isPathConnected_iff {f : X → Y} (hf : Inducing f) : IsPathConnected F ↔ IsPathConnected (f '' F) := by - refine ⟨fun hF ↦ hF.image hf.continuous, fun hF ↦ ?_⟩ - simp? [isPathConnected_iff] at hF ⊢ says - simp only [isPathConnected_iff, image_nonempty, mem_image, forall_exists_index, - and_imp, forall_apply_eq_imp_iff₂] at hF ⊢ - refine ⟨hF.1, fun x hx y hy ↦ ?_⟩ - rcases hF.2 x hx y hy with ⟨γ, hγ⟩ - choose γ' hγ' hγγ' using hγ - have key₁ : Inseparable x (γ' 0) := by rw [← hf.inseparable_iff, hγγ' 0, γ.source] - have key₂ : Inseparable (γ' 1) y := by rw [← hf.inseparable_iff, hγγ' 1, γ.target] - refine key₁.joinedIn hx (hγ' 0) |>.trans ⟨⟨⟨γ', ?_⟩, rfl, rfl⟩, hγ'⟩ |>.trans - (key₂.joinedIn (hγ' 1) hy) - simpa [hf.continuous_iff] using γ.continuous.congr fun t ↦ (hγγ' t).symm + simp only [IsPathConnected, forall_mem_image, exists_mem_image] + refine exists_congr fun x ↦ and_congr_right fun hx ↦ forall₂_congr fun y hy ↦ ?_ + rw [hf.joinedIn_image hx hy] /-- If `h : X → Y` is a homeomorphism, `h(s)` is path-connected iff `s` is. -/ @[simp] @@ -922,10 +936,7 @@ theorem IsPathConnected.union {U V : Set X} (hU : IsPathConnected U) (hV : IsPat ambient type `U` (when `U` contains `W`). -/ theorem IsPathConnected.preimage_coe {U W : Set X} (hW : IsPathConnected W) (hWU : W ⊆ U) : IsPathConnected (((↑) : U → X) ⁻¹' W) := by - rcases hW with ⟨x, x_in, hx⟩ - use ⟨x, hWU x_in⟩, by simp [x_in] - rintro ⟨y, hyU⟩ hyW - exact ⟨(hx hyW).joined_subtype.somePath.map (continuous_inclusion hWU), by simp⟩ + rwa [inducing_subtype_val.isPathConnected_iff, Subtype.image_preimage_val, inter_eq_right.2 hWU] theorem IsPathConnected.exists_path_through_family {n : ℕ} {s : Set X} (h : IsPathConnected s) (p : Fin (n + 1) → X) (hp : ∀ i, p i ∈ s) : @@ -997,6 +1008,7 @@ theorem IsPathConnected.exists_path_through_family' {n : ℕ} /-- A topological space is path-connected if it is non-empty and every two points can be joined by a continuous path. -/ +@[mk_iff] class PathConnectedSpace (X : Type*) [TopologicalSpace X] : Prop where /-- A path-connected space must be nonempty. -/ nonempty : Nonempty X @@ -1025,31 +1037,12 @@ def somePath (x y : X) : Path x y := end PathConnectedSpace -theorem isPathConnected_iff_pathConnectedSpace : IsPathConnected F ↔ PathConnectedSpace F := by - rw [isPathConnected_iff] - constructor - · rintro ⟨⟨x, x_in⟩, h⟩ - refine ⟨⟨⟨x, x_in⟩⟩, ?_⟩ - rintro ⟨y, y_in⟩ ⟨z, z_in⟩ - have H := h y y_in z z_in - rwa [joinedIn_iff_joined y_in z_in] at H - · rintro ⟨⟨x, x_in⟩, H⟩ - refine ⟨⟨x, x_in⟩, fun y y_in z z_in => ?_⟩ - rw [joinedIn_iff_joined y_in z_in] - apply H - theorem pathConnectedSpace_iff_univ : PathConnectedSpace X ↔ IsPathConnected (univ : Set X) := by - constructor - · intro h - haveI := @PathConnectedSpace.nonempty X _ _ - inhabit X - refine ⟨default, mem_univ _, ?_⟩ - intros y _hy - simpa using PathConnectedSpace.joined default y - · intro h - have h' := h.joinedIn - cases' h with x h - exact ⟨⟨x⟩, by simpa using h'⟩ + simp [pathConnectedSpace_iff, isPathConnected_iff, nonempty_iff_univ_nonempty] + +theorem isPathConnected_iff_pathConnectedSpace : IsPathConnected F ↔ PathConnectedSpace F := by + rw [pathConnectedSpace_iff_univ, inducing_subtype_val.isPathConnected_iff, image_univ, + Subtype.range_val_subtype, setOf_mem_eq] theorem isPathConnected_univ [PathConnectedSpace X] : IsPathConnected (univ : Set X) := pathConnectedSpace_iff_univ.mp inferInstance @@ -1111,6 +1104,7 @@ end PathConnectedSpace /-! ### Locally path connected spaces -/ +section LocPathConnectedSpace /-- A topological space is locally path connected, at every point, path connected neighborhoods form a neighborhood basis. -/ @@ -1120,66 +1114,70 @@ class LocPathConnectedSpace (X : Type*) [TopologicalSpace X] : Prop where export LocPathConnectedSpace (path_connected_basis) -theorem locPathConnected_of_bases {p : ι → Prop} {s : X → ι → Set X} - (h : ∀ x, (𝓝 x).HasBasis p (s x)) (h' : ∀ x i, p i → IsPathConnected (s x i)) : - LocPathConnectedSpace X := by - constructor - intro x - apply (h x).to_hasBasis - · intro i pi - exact ⟨s x i, ⟨(h x).mem_of_mem pi, h' x i pi⟩, by rfl⟩ - · rintro U ⟨U_in, _hU⟩ - rcases (h x).mem_iff.mp U_in with ⟨i, pi, hi⟩ - tauto - -theorem pathConnectedSpace_iff_connectedSpace [LocPathConnectedSpace X] : - PathConnectedSpace X ↔ ConnectedSpace X := by - constructor - · intro h - infer_instance - · intro hX - rw [pathConnectedSpace_iff_eq] - use Classical.arbitrary X - refine IsClopen.eq_univ ⟨?_, ?_⟩ (by simp) - · rw [isClosed_iff_nhds] - intro y H - rcases (path_connected_basis y).ex_mem with ⟨U, ⟨U_in, hU⟩⟩ - rcases H U U_in with ⟨z, hz, hz'⟩ - exact (hU.joinedIn z hz y <| mem_of_mem_nhds U_in).joined.mem_pathComponent hz' - · rw [isOpen_iff_mem_nhds] - intro y y_in - rcases (path_connected_basis y).ex_mem with ⟨U, ⟨U_in, hU⟩⟩ - apply mem_of_superset U_in - rw [← pathComponent_congr y_in] - exact hU.subset_pathComponent (mem_of_mem_nhds U_in) - -theorem pathConnected_subset_basis [LocPathConnectedSpace X] {U : Set X} (h : IsOpen U) - (hx : x ∈ U) : (𝓝 x).HasBasis (fun s : Set X => s ∈ 𝓝 x ∧ IsPathConnected s ∧ s ⊆ U) id := +theorem LocPathConnectedSpace.of_bases {p : X → ι → Prop} {s : X → ι → Set X} + (h : ∀ x, (𝓝 x).HasBasis (p x) (s x)) (h' : ∀ x i, p x i → IsPathConnected (s x i)) : + LocPathConnectedSpace X where + path_connected_basis x := by + rw [hasBasis_self] + intro t ht + rcases (h x).mem_iff.mp ht with ⟨i, hpi, hi⟩ + exact ⟨s x i, (h x).mem_of_mem hpi, h' x i hpi, hi⟩ + +@[deprecated (since := "2024-10-16")] +alias locPathConnected_of_bases := LocPathConnectedSpace.of_bases + +variable [LocPathConnectedSpace X] + +/-- In a locally path connected space, each path component is an open set. -/ +protected theorem IsOpen.pathComponent (x : X) : IsOpen (pathComponent x) := by + rw [isOpen_iff_mem_nhds] + intro y hxy + rcases (path_connected_basis y).ex_mem with ⟨V, hVy, hVc⟩ + filter_upwards [hVy] with z hz + exact hxy.out.trans (hVc.joinedIn _ (mem_of_mem_nhds hVy) _ hz).joined + +/-- In a locally path connected space, each path component is a closed set. -/ +protected theorem IsClosed.pathComponent (x : X) : IsClosed (pathComponent x) := by + rw [← isOpen_compl_iff, isOpen_iff_mem_nhds] + intro y hxy + rcases (path_connected_basis y).ex_mem with ⟨V, hVy, hVc⟩ + filter_upwards [hVy] with z hz hxz + exact hxy <| hxz.trans (hVc.joinedIn _ hz _ (mem_of_mem_nhds hVy)).joined + +/-- In a locally path connected space, each path component is a clopen set. -/ +protected theorem IsClopen.pathComponent (x : X) : IsClopen (pathComponent x) := + ⟨.pathComponent x, .pathComponent x⟩ + +theorem pathConnectedSpace_iff_connectedSpace : PathConnectedSpace X ↔ ConnectedSpace X := by + refine ⟨fun _ ↦ inferInstance, fun h ↦ ⟨inferInstance, fun x y ↦ ?_⟩⟩ + rw [← mem_pathComponent_iff, (IsClopen.pathComponent _).eq_univ] <;> simp + +theorem pathComponent_eq_connectedComponent (x : X) : pathComponent x = connectedComponent x := + (pathComponent_subset_component x).antisymm <| + (IsClopen.pathComponent x).connectedComponent_subset (mem_pathComponent_self _) + +theorem pathConnected_subset_basis {U : Set X} (h : IsOpen U) (hx : x ∈ U) : + (𝓝 x).HasBasis (fun s : Set X => s ∈ 𝓝 x ∧ IsPathConnected s ∧ s ⊆ U) id := (path_connected_basis x).hasBasis_self_subset (IsOpen.mem_nhds h hx) -theorem locPathConnected_of_isOpen [LocPathConnectedSpace X] {U : Set X} (h : IsOpen U) : - LocPathConnectedSpace U := - ⟨by - rintro ⟨x, x_in⟩ - rw [nhds_subtype_eq_comap] - constructor - intro V - rw [(HasBasis.comap ((↑) : U → X) (pathConnected_subset_basis h x_in)).mem_iff] - constructor - · rintro ⟨W, ⟨W_in, hW, hWU⟩, hWV⟩ - exact ⟨Subtype.val ⁻¹' W, ⟨⟨preimage_mem_comap W_in, hW.preimage_coe hWU⟩, hWV⟩⟩ - · rintro ⟨W, ⟨W_in, hW⟩, hWV⟩ - refine - ⟨(↑) '' W, - ⟨Filter.image_coe_mem_of_mem_comap (IsOpen.mem_nhds h x_in) W_in, - hW.image continuous_subtype_val, Subtype.coe_image_subset U W⟩, - ?_⟩ - rintro x ⟨y, ⟨y_in, hy⟩⟩ - rw [← Subtype.coe_injective hy] - tauto⟩ - -theorem IsOpen.isConnected_iff_isPathConnected [LocPathConnectedSpace X] {U : Set X} - (U_op : IsOpen U) : IsPathConnected U ↔ IsConnected U := by +theorem OpenEmbedding.locPathConnectedSpace {e : Y → X} (he : OpenEmbedding e) : + LocPathConnectedSpace Y := + have (y : Y) : + (𝓝 y).HasBasis (fun s ↦ s ∈ 𝓝 (e y) ∧ IsPathConnected s ∧ s ⊆ range e) (e ⁻¹' ·) := + he.basis_nhds <| pathConnected_subset_basis he.isOpen_range (mem_range_self _) + .of_bases this fun x s ⟨_, hs, hse⟩ ↦ by + rwa [he.isPathConnected_iff, image_preimage_eq_of_subset hse] + +theorem IsOpen.locPathConnectedSpace {U : Set X} (h : IsOpen U) : LocPathConnectedSpace U := + (openEmbedding_subtype_val h).locPathConnectedSpace + +@[deprecated (since := "2024-10-17")] +alias locPathConnected_of_isOpen := IsOpen.locPathConnectedSpace + +theorem IsOpen.isConnected_iff_isPathConnected {U : Set X} (U_op : IsOpen U) : + IsConnected U ↔ IsPathConnected U := by rw [isConnected_iff_connectedSpace, isPathConnected_iff_pathConnectedSpace] - haveI := locPathConnected_of_isOpen U_op - exact pathConnectedSpace_iff_connectedSpace + haveI := U_op.locPathConnectedSpace + exact pathConnectedSpace_iff_connectedSpace.symm + +end LocPathConnectedSpace From 319d7a37f0e813c9603b8df05901330e49aaf0eb Mon Sep 17 00:00:00 2001 From: FR Date: Fri, 18 Oct 2024 06:38:39 +0000 Subject: [PATCH 125/505] chore: move `NNRat` APIs earlier (#17269) No need to import bundled ordered algebraic typeclasses in `ring` tactic now. --- Archive/Imo/Imo1998Q2.lean | 2 +- Archive/Imo/Imo2013Q1.lean | 7 ++- .../ContinuedFractions/ConvergentsEquiv.lean | 1 + Mathlib/Algebra/Field/Rat.lean | 52 +++++++++++++++++-- Mathlib/Algebra/Order/Field/Rat.lean | 22 -------- Mathlib/Algebra/Order/Module/Rat.lean | 1 + Mathlib/Algebra/Order/Star/Basic.lean | 4 +- Mathlib/Data/FP/Basic.lean | 2 +- Mathlib/Data/Finset/Density.lean | 1 + Mathlib/Data/NNRat/BigOperators.lean | 1 + Mathlib/Data/NNRat/Defs.lean | 11 +++- Mathlib/Data/NNRat/Lemmas.lean | 7 --- Mathlib/Data/Nat/BitIndices.lean | 6 +-- Mathlib/Data/Rat/Cast/Defs.lean | 7 +-- Mathlib/Data/Rat/Cast/Lemmas.lean | 2 +- Mathlib/Data/Rat/Cast/Order.lean | 2 +- Mathlib/Data/Rat/Star.lean | 4 +- Mathlib/Data/Real/Basic.lean | 1 + Mathlib/NumberTheory/Fermat.lean | 1 + Mathlib/Tactic/Polyrith.lean | 1 - Mathlib/Tactic/Ring/Basic.lean | 3 +- test/cancel_denoms.lean | 1 + test/linear_combination'.lean | 1 + test/linear_combination.lean | 1 + test/module.lean | 1 + test/ring.lean | 1 + 26 files changed, 88 insertions(+), 55 deletions(-) diff --git a/Archive/Imo/Imo1998Q2.lean b/Archive/Imo/Imo1998Q2.lean index 367aec35fd5fc..db7ae35345496 100644 --- a/Archive/Imo/Imo1998Q2.lean +++ b/Archive/Imo/Imo1998Q2.lean @@ -5,7 +5,7 @@ Authors: Oliver Nash -/ import Mathlib.Algebra.Order.BigOperators.Group.Finset import Mathlib.Algebra.Order.Field.Basic -import Mathlib.Algebra.Ring.Int +import Mathlib.Algebra.Order.Field.Rat import Mathlib.GroupTheory.GroupAction.Ring import Mathlib.Tactic.NoncommRing import Mathlib.Tactic.Ring diff --git a/Archive/Imo/Imo2013Q1.lean b/Archive/Imo/Imo2013Q1.lean index 42b43f937926c..295886d7e18ea 100644 --- a/Archive/Imo/Imo2013Q1.lean +++ b/Archive/Imo/Imo2013Q1.lean @@ -3,12 +3,11 @@ Copyright (c) 2021 David Renshaw. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: David Renshaw -/ -import Mathlib.Algebra.BigOperators.Pi -import Mathlib.Algebra.Order.Ring.Abs -import Mathlib.Data.PNat.Basic -import Mathlib.Tactic.Ring +import Mathlib.Algebra.BigOperators.Group.Finset +import Mathlib.Algebra.Order.Field.Rat import Mathlib.Tactic.FieldSimp import Mathlib.Tactic.Positivity.Basic +import Mathlib.Tactic.Ring /-! # IMO 2013 Q1 diff --git a/Mathlib/Algebra/ContinuedFractions/ConvergentsEquiv.lean b/Mathlib/Algebra/ContinuedFractions/ConvergentsEquiv.lean index 99e76fa64a5f5..fa5e59f20d0ad 100644 --- a/Mathlib/Algebra/ContinuedFractions/ConvergentsEquiv.lean +++ b/Mathlib/Algebra/ContinuedFractions/ConvergentsEquiv.lean @@ -3,6 +3,7 @@ Copyright (c) 2020 Kevin Kappelmann. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kevin Kappelmann -/ +import Mathlib.Algebra.Order.Field.Defs import Mathlib.Algebra.ContinuedFractions.ContinuantsRecurrence import Mathlib.Algebra.ContinuedFractions.TerminatedStable import Mathlib.Tactic.FieldSimp diff --git a/Mathlib/Algebra/Field/Rat.lean b/Mathlib/Algebra/Field/Rat.lean index ee2d2e48680d1..ad60ac2ab4bb0 100644 --- a/Mathlib/Algebra/Field/Rat.lean +++ b/Mathlib/Algebra/Field/Rat.lean @@ -13,11 +13,6 @@ This file contains the field instance on the rational numbers. See note [foundational algebra order theory]. -## TODO - -Move the `Semifield ℚ≥0` instance here. This will involve proving it by hand rather than relying on -the `Nonneg` machinery. - ## Tags rat, rationals, field, ℚ, numerator, denominator, num, denom @@ -44,4 +39,51 @@ These also prevent non-computable instances being used to construct these instan instance instDivisionRing : DivisionRing ℚ := inferInstance +protected lemma inv_nonneg {a : ℚ} (ha : 0 ≤ a) : 0 ≤ a⁻¹ := by + rw [inv_def'] + exact divInt_nonneg (Int.ofNat_nonneg a.den) (num_nonneg.mpr ha) + +protected lemma div_nonneg {a b : ℚ} (ha : 0 ≤ a) (hb : 0 ≤ b) : 0 ≤ a / b := + mul_nonneg ha (Rat.inv_nonneg hb) + end Rat + +namespace NNRat + +instance instInv : Inv ℚ≥0 where + inv x := ⟨x⁻¹, Rat.inv_nonneg x.2⟩ + +instance instDiv : Div ℚ≥0 where + div x y := ⟨x / y, Rat.div_nonneg x.2 y.2⟩ + +@[simp, norm_cast] lemma coe_inv (q : ℚ≥0) : ((q⁻¹ : ℚ≥0) : ℚ) = (q : ℚ)⁻¹ := rfl +@[simp, norm_cast] lemma coe_div (p q : ℚ≥0) : ((p / q : ℚ≥0) : ℚ) = p / q := rfl + +lemma inv_def (q : ℚ≥0) : q⁻¹ = divNat q.den q.num := by ext; simp [Rat.inv_def', num_coe, den_coe] +lemma div_def (p q : ℚ≥0) : p / q = divNat (p.num * q.den) (p.den * q.num) := by + ext; simp [Rat.div_def', num_coe, den_coe] + +lemma num_inv_of_ne_zero {q : ℚ≥0} (hq : q ≠ 0) : q⁻¹.num = q.den := by + rw [inv_def, divNat, num, coe_mk, Rat.divInt_ofNat, ← Rat.mk_eq_mkRat _ _ (num_ne_zero.mpr hq), + Int.natAbs_ofNat] + simpa using q.coprime_num_den.symm + +lemma den_inv_of_ne_zero {q : ℚ≥0} (hq : q ≠ 0) : q⁻¹.den = q.num := by + rw [inv_def, divNat, den, coe_mk, Rat.divInt_ofNat, ← Rat.mk_eq_mkRat _ _ (num_ne_zero.mpr hq)] + simpa using q.coprime_num_den.symm + +@[simp] +lemma num_div_den (q : ℚ≥0) : (q.num : ℚ≥0) / q.den = q := by + ext1 + rw [coe_div, coe_natCast, coe_natCast, num, ← Int.cast_natCast] + exact (cast_def _).symm + +instance instSemifield : Semifield ℚ≥0 where + __ := instNNRatCommSemiring + inv_zero := by ext; simp + mul_inv_cancel q h := by ext; simp [h] + nnratCast_def q := q.num_div_den.symm + nnqsmul q a := q * a + nnqsmul_def q a := rfl + +end NNRat diff --git a/Mathlib/Algebra/Order/Field/Rat.lean b/Mathlib/Algebra/Order/Field/Rat.lean index 228df54d19272..7b6ba3bfd819a 100644 --- a/Mathlib/Algebra/Order/Field/Rat.lean +++ b/Mathlib/Algebra/Order/Field/Rat.lean @@ -32,25 +32,3 @@ end Rat -- instances for performance deriving instance CanonicallyLinearOrderedSemifield, LinearOrderedSemifield, LinearOrderedCommGroupWithZero for NNRat - -/-! ### Miscellaneous lemmas -/ - -namespace NNRat - -@[simp, norm_cast] lemma coe_inv (q : ℚ≥0) : ((q⁻¹ : ℚ≥0) : ℚ) = (q : ℚ)⁻¹ := rfl -@[simp, norm_cast] lemma coe_div (p q : ℚ≥0) : ((p / q : ℚ≥0) : ℚ) = p / q := rfl - -lemma inv_def (q : ℚ≥0) : q⁻¹ = divNat q.den q.num := by ext; simp [Rat.inv_def', num_coe, den_coe] -lemma div_def (p q : ℚ≥0) : p / q = divNat (p.num * q.den) (p.den * q.num) := by - ext; simp [Rat.div_def', num_coe, den_coe] - -lemma num_inv_of_ne_zero {q : ℚ≥0} (hq : q ≠ 0) : q⁻¹.num = q.den := by - rw [inv_def, divNat, num, coe_mk, Rat.divInt_ofNat, ← Rat.mk_eq_mkRat _ _ (num_ne_zero.mpr hq), - Int.natAbs_ofNat] - simpa using q.coprime_num_den.symm - -lemma den_inv_of_ne_zero {q : ℚ≥0} (hq : q ≠ 0) : q⁻¹.den = q.num := by - rw [inv_def, divNat, den, coe_mk, Rat.divInt_ofNat, ← Rat.mk_eq_mkRat _ _ (num_ne_zero.mpr hq)] - simpa using q.coprime_num_den.symm - -end NNRat diff --git a/Mathlib/Algebra/Order/Module/Rat.lean b/Mathlib/Algebra/Order/Module/Rat.lean index 288226c006963..21b376b2c7771 100644 --- a/Mathlib/Algebra/Order/Module/Rat.lean +++ b/Mathlib/Algebra/Order/Module/Rat.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ import Mathlib.Algebra.Order.Module.Defs +import Mathlib.Data.NNRat.Lemmas import Mathlib.Data.Rat.Cast.Order /-! diff --git a/Mathlib/Algebra/Order/Star/Basic.lean b/Mathlib/Algebra/Order/Star/Basic.lean index bcda0ea6851ab..6370c2c91944e 100644 --- a/Mathlib/Algebra/Order/Star/Basic.lean +++ b/Mathlib/Algebra/Order/Star/Basic.lean @@ -3,10 +3,10 @@ Copyright (c) 2023 Kim Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison -/ -import Mathlib.Algebra.Group.Submonoid.Operations +import Mathlib.Algebra.Order.Group.Defs +import Mathlib.Algebra.Order.Group.Nat import Mathlib.Algebra.Star.SelfAdjoint import Mathlib.Algebra.Star.StarRingHom -import Mathlib.Algebra.Regular.Basic import Mathlib.Tactic.ContinuousFunctionalCalculus /-! # Star ordered rings diff --git a/Mathlib/Data/FP/Basic.lean b/Mathlib/Data/FP/Basic.lean index c854edb12daef..703af9e82aadb 100644 --- a/Mathlib/Data/FP/Basic.lean +++ b/Mathlib/Data/FP/Basic.lean @@ -85,7 +85,7 @@ theorem Float.Zero.valid : ValidFinite emin 0 := ring_nf rw [mul_comm] assumption - le_trans C.precMax (Nat.le_mul_of_pos_left _ two_pos), + le_trans C.precMax (Nat.le_mul_of_pos_left _ Nat.zero_lt_two), by (rw [max_eq_right]; simp [sub_eq_add_neg, Int.ofNat_zero_le])⟩ @[nolint docBlame] diff --git a/Mathlib/Data/Finset/Density.lean b/Mathlib/Data/Finset/Density.lean index 002513c40e808..08b5292155a45 100644 --- a/Mathlib/Data/Finset/Density.lean +++ b/Mathlib/Data/Finset/Density.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ import Mathlib.Algebra.Order.Field.Basic +import Mathlib.Algebra.Order.Field.Rat import Mathlib.Data.Fintype.Card import Mathlib.Data.NNRat.Order import Mathlib.Data.Rat.Cast.CharZero diff --git a/Mathlib/Data/NNRat/BigOperators.lean b/Mathlib/Data/NNRat/BigOperators.lean index e49a23c8a08e6..755e9be98bc81 100644 --- a/Mathlib/Data/NNRat/BigOperators.lean +++ b/Mathlib/Data/NNRat/BigOperators.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies, Bhavik Mehta -/ import Mathlib.Algebra.Order.BigOperators.Ring.Finset +import Mathlib.Algebra.Order.Ring.Rat import Mathlib.Data.NNRat.Defs /-! # Casting lemmas for non-negative rational numbers involving sums and products diff --git a/Mathlib/Data/NNRat/Defs.lean b/Mathlib/Data/NNRat/Defs.lean index 5bda9ba5b3826..3f411698b84cd 100644 --- a/Mathlib/Data/NNRat/Defs.lean +++ b/Mathlib/Data/NNRat/Defs.lean @@ -14,7 +14,8 @@ import Mathlib.Algebra.Ring.Rat This file defines the nonnegative rationals as a subtype of `Rat` and provides its basic algebraic order structure. -Note that `NNRat` is not declared as a `Field` here. See `Data.NNRat.Lemmas` for that instance. +Note that `NNRat` is not declared as a `Semifield` here. See `Mathlib.Algebra.Field.Rat` for that +instance. We also define an instance `CanLift ℚ ℚ≥0`. This instance can be used by the `lift` tactic to replace `x : ℚ` and `hx : 0 ≤ x` in the proof context with `x : ℚ≥0` while replacing all occurrences @@ -63,8 +64,16 @@ namespace NNRat variable {p q : ℚ≥0} +instance instNontrivial : Nontrivial ℚ≥0 where exists_pair_ne := ⟨1, 0, by decide⟩ +instance instOrderBot : OrderBot ℚ≥0 where + bot := 0 + bot_le q := q.2 + @[simp] lemma val_eq_cast (q : ℚ≥0) : q.1 = q := rfl +instance instCharZero : CharZero ℚ≥0 where + cast_injective a b hab := by simpa using congr_arg num hab + instance canLift : CanLift ℚ ℚ≥0 (↑) fun q ↦ 0 ≤ q where prf q hq := ⟨⟨q, hq⟩, rfl⟩ diff --git a/Mathlib/Data/NNRat/Lemmas.lean b/Mathlib/Data/NNRat/Lemmas.lean index 4e1ee467939eb..91aadd51015c4 100644 --- a/Mathlib/Data/NNRat/Lemmas.lean +++ b/Mathlib/Data/NNRat/Lemmas.lean @@ -61,13 +61,6 @@ namespace NNRat variable {p q : ℚ≥0} -@[simp] -lemma num_div_den (q : ℚ≥0) : (q.num : ℚ≥0) / q.den = q := by - ext : 1 - rw [coe_div, coe_natCast, coe_natCast, num, ← Int.cast_natCast, - Int.natAbs_of_nonneg (Rat.num_nonneg.2 q.cast_nonneg)] - exact Rat.num_div_den q - /-- A recursor for nonnegative rationals in terms of numerators and denominators. -/ protected def rec {α : ℚ≥0 → Sort*} (h : ∀ m n : ℕ, α (m / n)) (q : ℚ≥0) : α q := by rw [← num_div_den q]; apply h diff --git a/Mathlib/Data/Nat/BitIndices.lean b/Mathlib/Data/Nat/BitIndices.lean index 79bd2eb583ef4..bd77462f10032 100644 --- a/Mathlib/Data/Nat/BitIndices.lean +++ b/Mathlib/Data/Nat/BitIndices.lean @@ -3,12 +3,12 @@ Copyright (c) 2024 Peter Nelson. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Peter Nelson -/ -import Mathlib.Data.List.Sort -import Mathlib.Data.Nat.Bitwise import Mathlib.Algebra.BigOperators.Ring.List import Mathlib.Algebra.Order.BigOperators.Group.List import Mathlib.Algebra.Order.Star.Basic -import Mathlib.Algebra.Order.Sub.Defs +import Mathlib.Algebra.Order.Sub.Basic +import Mathlib.Data.List.Sort +import Mathlib.Data.Nat.Bitwise /-! # Bit Indices diff --git a/Mathlib/Data/Rat/Cast/Defs.lean b/Mathlib/Data/Rat/Cast/Defs.lean index 739dbc17cb0ce..5458a333b2468 100644 --- a/Mathlib/Data/Rat/Cast/Defs.lean +++ b/Mathlib/Data/Rat/Cast/Defs.lean @@ -3,13 +3,12 @@ Copyright (c) 2019 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro -/ +import Mathlib.Algebra.Field.Basic import Mathlib.Algebra.Field.Rat import Mathlib.Algebra.Group.Commute.Basic -import Mathlib.Algebra.GroupWithZero.Units.Lemmas -import Mathlib.Algebra.Order.Field.Rat import Mathlib.Data.Int.Cast.Lemmas -import Mathlib.Data.NNRat.Lemmas import Mathlib.Data.Rat.Lemmas +import Mathlib.Order.Nat /-! # Casts for Rational Numbers @@ -24,6 +23,8 @@ casting lemmas showing the well-behavedness of this injection. rat, rationals, field, ℚ, numerator, denominator, num, denom, cast, coercion, casting -/ +assert_not_exists OrderedAddCommMonoid + variable {F ι α β : Type*} namespace NNRat diff --git a/Mathlib/Data/Rat/Cast/Lemmas.lean b/Mathlib/Data/Rat/Cast/Lemmas.lean index 4089d6b56d2e4..f61ef4b456824 100644 --- a/Mathlib/Data/Rat/Cast/Lemmas.lean +++ b/Mathlib/Data/Rat/Cast/Lemmas.lean @@ -3,7 +3,7 @@ Copyright (c) 2019 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro -/ -import Mathlib.Algebra.Field.Basic +import Mathlib.Algebra.Order.Nonneg.Field import Mathlib.Data.Rat.Cast.Defs import Mathlib.Tactic.Positivity.Basic diff --git a/Mathlib/Data/Rat/Cast/Order.lean b/Mathlib/Data/Rat/Cast/Order.lean index 2a1d627e9d6b0..c85aa1ca681f0 100644 --- a/Mathlib/Data/Rat/Cast/Order.lean +++ b/Mathlib/Data/Rat/Cast/Order.lean @@ -3,7 +3,7 @@ Copyright (c) 2019 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro -/ -import Mathlib.Algebra.Order.Ring.Rat +import Mathlib.Algebra.Order.Field.Rat import Mathlib.Data.Rat.Cast.CharZero import Mathlib.Tactic.Positivity.Core import Mathlib.Algebra.Order.Field.Basic diff --git a/Mathlib/Data/Rat/Star.lean b/Mathlib/Data/Rat/Star.lean index 8f4b38afcab0e..1e69ef6c77136 100644 --- a/Mathlib/Data/Rat/Star.lean +++ b/Mathlib/Data/Rat/Star.lean @@ -4,10 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jireh Loreaux, Yaël Dillies -/ import Mathlib.Algebra.GroupWithZero.Commute +import Mathlib.Algebra.Order.Field.Rat +import Mathlib.Algebra.Order.Monoid.Submonoid import Mathlib.Algebra.Order.Ring.Abs import Mathlib.Algebra.Order.Star.Basic -import Mathlib.Data.NNRat.Lemmas -import Mathlib.Algebra.Order.Monoid.Submonoid import Mathlib.Tactic.FieldSimp /-! diff --git a/Mathlib/Data/Real/Basic.lean b/Mathlib/Data/Real/Basic.lean index 9efe0fb5d9ae3..3465de02f2a9a 100644 --- a/Mathlib/Data/Real/Basic.lean +++ b/Mathlib/Data/Real/Basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Floris van Doorn -/ import Mathlib.Algebra.Order.CauSeq.Completion +import Mathlib.Algebra.Order.Field.Rat /-! # Real numbers from Cauchy sequences diff --git a/Mathlib/NumberTheory/Fermat.lean b/Mathlib/NumberTheory/Fermat.lean index 848dc49be4810..c90f62582cae8 100644 --- a/Mathlib/NumberTheory/Fermat.lean +++ b/Mathlib/NumberTheory/Fermat.lean @@ -5,6 +5,7 @@ Authors: Moritz Firsching -/ import Mathlib.Algebra.BigOperators.Group.Finset import Mathlib.Algebra.Order.Ring.Basic +import Mathlib.Algebra.Order.Ring.Nat import Mathlib.Algebra.Order.Star.Basic import Mathlib.Data.Nat.Prime.Defs import Mathlib.Tactic.Ring.RingNF diff --git a/Mathlib/Tactic/Polyrith.lean b/Mathlib/Tactic/Polyrith.lean index 1b3ef56ab08c8..b6da0ce70e38b 100644 --- a/Mathlib/Tactic/Polyrith.lean +++ b/Mathlib/Tactic/Polyrith.lean @@ -3,7 +3,6 @@ Copyright (c) 2022 Dhruv Bhatia. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Dhruv Bhatia, Eric Wieser, Mario Carneiro -/ -import Mathlib.Algebra.Order.Field.Rat import Mathlib.Tactic.LinearCombination /-! diff --git a/Mathlib/Tactic/Ring/Basic.lean b/Mathlib/Tactic/Ring/Basic.lean index c5ac7fd426cd8..b25109521ba02 100644 --- a/Mathlib/Tactic/Ring/Basic.lean +++ b/Mathlib/Tactic/Ring/Basic.lean @@ -3,7 +3,6 @@ Copyright (c) 2018 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Aurélien Saue, Anne Baanen -/ -import Mathlib.Algebra.Order.Ring.Rat import Mathlib.Tactic.NormNum.Inv import Mathlib.Tactic.NormNum.Pow import Mathlib.Util.AtomM @@ -75,6 +74,8 @@ This feature wasn't needed yet, so it's not implemented yet. ring, semiring, exponent, power -/ +assert_not_exists OrderedAddCommMonoid + namespace Mathlib.Tactic namespace Ring diff --git a/test/cancel_denoms.lean b/test/cancel_denoms.lean index b717b67f56ee2..4d1001ecbcd6e 100644 --- a/test/cancel_denoms.lean +++ b/test/cancel_denoms.lean @@ -1,3 +1,4 @@ +import Mathlib.Algebra.Order.Field.Rat import Mathlib.Tactic.CancelDenoms import Mathlib.Tactic.Ring diff --git a/test/linear_combination'.lean b/test/linear_combination'.lean index 14516504da879..36056748d622d 100644 --- a/test/linear_combination'.lean +++ b/test/linear_combination'.lean @@ -1,3 +1,4 @@ +import Mathlib.Algebra.Order.Field.Defs import Mathlib.Tactic.LinearCombination' import Mathlib.Tactic.Linarith diff --git a/test/linear_combination.lean b/test/linear_combination.lean index 9f652a8186a63..6e8b23a79630f 100644 --- a/test/linear_combination.lean +++ b/test/linear_combination.lean @@ -1,3 +1,4 @@ +import Mathlib.Algebra.Order.Field.Defs import Mathlib.Tactic.LinearCombination import Mathlib.Tactic.Linarith diff --git a/test/module.lean b/test/module.lean index ee59dd02d0d9f..b3ec8f57b5789 100644 --- a/test/module.lean +++ b/test/module.lean @@ -3,6 +3,7 @@ Copyright (c) 2024 Heather Macbeth. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Heather Macbeth -/ +import Mathlib.Algebra.Order.Field.Defs import Mathlib.Tactic.FieldSimp import Mathlib.Tactic.LinearCombination import Mathlib.Tactic.Module diff --git a/test/ring.lean b/test/ring.lean index 5abd490bfa500..ffdfe755e4dec 100644 --- a/test/ring.lean +++ b/test/ring.lean @@ -1,3 +1,4 @@ +import Mathlib.Algebra.Order.Field.Defs import Mathlib.Tactic.FieldSimp import Mathlib.Tactic.Ring From 9b16c5d7fb54a2f624ce50e9b4f9de1e2a26b963 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 18 Oct 2024 07:23:26 +0000 Subject: [PATCH 126/505] chore(Topology/Algebra/Order): move 2 files (#17879) These two files don't deal with algebra (groups, monoids etc); move them to `Topology/Order` instead. --- Mathlib.lean | 4 ++-- Mathlib/Analysis/Calculus/LocalExtr/Rolle.lean | 2 +- Mathlib/Topology/MetricSpace/Bounded.lean | 2 +- Mathlib/Topology/MetricSpace/ProperSpace/Lemmas.lean | 2 +- Mathlib/Topology/{Algebra => }/Order/Compact.lean | 0 Mathlib/Topology/{Algebra => }/Order/Rolle.lean | 2 +- test/LibrarySearch/IsCompact.lean | 2 +- 7 files changed, 7 insertions(+), 7 deletions(-) rename Mathlib/Topology/{Algebra => }/Order/Compact.lean (100%) rename Mathlib/Topology/{Algebra => }/Order/Rolle.lean (98%) diff --git a/Mathlib.lean b/Mathlib.lean index 9a66b6063fc59..6b0d74489e5e7 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4561,12 +4561,10 @@ import Mathlib.Topology.Algebra.Nonarchimedean.Completion import Mathlib.Topology.Algebra.Nonarchimedean.TotallyDisconnected import Mathlib.Topology.Algebra.OpenSubgroup import Mathlib.Topology.Algebra.Order.Archimedean -import Mathlib.Topology.Algebra.Order.Compact import Mathlib.Topology.Algebra.Order.Field import Mathlib.Topology.Algebra.Order.Floor import Mathlib.Topology.Algebra.Order.Group import Mathlib.Topology.Algebra.Order.LiminfLimsup -import Mathlib.Topology.Algebra.Order.Rolle import Mathlib.Topology.Algebra.Order.UpperLower import Mathlib.Topology.Algebra.Polynomial import Mathlib.Topology.Algebra.PontryaginDual @@ -4823,6 +4821,7 @@ import Mathlib.Topology.Order.Bornology import Mathlib.Topology.Order.Bounded import Mathlib.Topology.Order.Category.AlexDisc import Mathlib.Topology.Order.Category.FrameAdjunction +import Mathlib.Topology.Order.Compact import Mathlib.Topology.Order.DenselyOrdered import Mathlib.Topology.Order.ExtendFrom import Mathlib.Topology.Order.ExtrClosure @@ -4846,6 +4845,7 @@ import Mathlib.Topology.Order.OrderClosed import Mathlib.Topology.Order.PartialSups import Mathlib.Topology.Order.Priestley import Mathlib.Topology.Order.ProjIcc +import Mathlib.Topology.Order.Rolle import Mathlib.Topology.Order.ScottTopology import Mathlib.Topology.Order.T5 import Mathlib.Topology.Order.UpperLowerSetTopology diff --git a/Mathlib/Analysis/Calculus/LocalExtr/Rolle.lean b/Mathlib/Analysis/Calculus/LocalExtr/Rolle.lean index bd4dace2f8094..ef3e8f191d380 100644 --- a/Mathlib/Analysis/Calculus/LocalExtr/Rolle.lean +++ b/Mathlib/Analysis/Calculus/LocalExtr/Rolle.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov, Anatole Dedecker -/ import Mathlib.Analysis.Calculus.LocalExtr.Basic -import Mathlib.Topology.Algebra.Order.Rolle +import Mathlib.Topology.Order.Rolle /-! # Rolle's Theorem diff --git a/Mathlib/Topology/MetricSpace/Bounded.lean b/Mathlib/Topology/MetricSpace/Bounded.lean index 5ce1a3fd60a5e..7bd2d18acad75 100644 --- a/Mathlib/Topology/MetricSpace/Bounded.lean +++ b/Mathlib/Topology/MetricSpace/Bounded.lean @@ -3,7 +3,7 @@ Copyright (c) 2015, 2017 Jeremy Avigad. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Robert Y. Lewis, Johannes Hölzl, Mario Carneiro, Sébastien Gouëzel -/ -import Mathlib.Topology.Algebra.Order.Compact +import Mathlib.Topology.Order.Compact import Mathlib.Topology.MetricSpace.ProperSpace import Mathlib.Topology.MetricSpace.Cauchy import Mathlib.Topology.EMetricSpace.Diam diff --git a/Mathlib/Topology/MetricSpace/ProperSpace/Lemmas.lean b/Mathlib/Topology/MetricSpace/ProperSpace/Lemmas.lean index d582d83e1a6e6..f75b63205c5b6 100644 --- a/Mathlib/Topology/MetricSpace/ProperSpace/Lemmas.lean +++ b/Mathlib/Topology/MetricSpace/ProperSpace/Lemmas.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ -import Mathlib.Topology.Algebra.Order.Compact +import Mathlib.Topology.Order.Compact import Mathlib.Topology.MetricSpace.ProperSpace import Mathlib.Topology.Order.IntermediateValue import Mathlib.Topology.Order.LocalExtr diff --git a/Mathlib/Topology/Algebra/Order/Compact.lean b/Mathlib/Topology/Order/Compact.lean similarity index 100% rename from Mathlib/Topology/Algebra/Order/Compact.lean rename to Mathlib/Topology/Order/Compact.lean diff --git a/Mathlib/Topology/Algebra/Order/Rolle.lean b/Mathlib/Topology/Order/Rolle.lean similarity index 98% rename from Mathlib/Topology/Algebra/Order/Rolle.lean rename to Mathlib/Topology/Order/Rolle.lean index 57a3118c4e740..31fa140f80c8a 100644 --- a/Mathlib/Topology/Algebra/Order/Rolle.lean +++ b/Mathlib/Topology/Order/Rolle.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import Mathlib.Topology.Order.ExtendFrom -import Mathlib.Topology.Algebra.Order.Compact +import Mathlib.Topology.Order.Compact import Mathlib.Topology.Order.LocalExtr import Mathlib.Topology.Order.T5 diff --git a/test/LibrarySearch/IsCompact.lean b/test/LibrarySearch/IsCompact.lean index 0428a0cc91ba1..0ac0af98ca6a4 100644 --- a/test/LibrarySearch/IsCompact.lean +++ b/test/LibrarySearch/IsCompact.lean @@ -1,5 +1,5 @@ import Mathlib.Topology.Instances.Real -import Mathlib.Topology.Algebra.Order.Compact +import Mathlib.Topology.Order.Compact -- TODO: uses sorry, but is hidden behind the `apply?` /-- warning: declaration uses 'sorry' -/ From 672f75b15914a0b4f9b9f6aa3108cb0950e1cf65 Mon Sep 17 00:00:00 2001 From: Jiang Jiedong Date: Fri, 18 Oct 2024 07:51:31 +0000 Subject: [PATCH 127/505] refactor: rename `IsLocalRingHom` to `IsLocalHom` (#17403) Replace the `IsLocalRingHom` with `IsLocalHom`. Add instances for `IsLocalHom` in `CommRingCat` to resolve the problem caused by #6045. For a more detailed explanation of the problem, see https://github.com/leanprover-community/mathlib4/pull/17403#issuecomment-2418322525. Further golf some proof related to `IsLocalHom` instance search. Co-authored-by: Jiang Jiedong <107380768+jjdishere@users.noreply.github.com> --- Mathlib/Algebra/Associated/Basic.lean | 2 +- .../Algebra/Category/Ring/Constructions.lean | 14 +++-- Mathlib/Algebra/Category/Ring/Instances.lean | 52 +++++++++++++--- Mathlib/Algebra/Group/Units/Equiv.lean | 8 ++- Mathlib/Algebra/Group/Units/Hom.lean | 59 ++++++++++++------- .../Algebra/GroupWithZero/Units/Lemmas.lean | 11 ++-- Mathlib/Algebra/Polynomial/Expand.lean | 7 ++- Mathlib/AlgebraicGeometry/AffineScheme.lean | 1 - .../GammaSpecAdjunction.lean | 2 - .../PrimeSpectrum/Basic.lean | 11 ++-- .../ProjectiveSpectrum/Scheme.lean | 2 +- Mathlib/AlgebraicGeometry/ResidueField.lean | 9 --- Mathlib/AlgebraicGeometry/Scheme.lean | 5 +- Mathlib/AlgebraicGeometry/Spec.lean | 6 +- Mathlib/AlgebraicGeometry/Stalk.lean | 14 ++--- Mathlib/AlgebraicGeometry/StructureSheaf.lean | 8 +-- Mathlib/FieldTheory/Separable.lean | 2 +- .../RingedSpace/LocallyRingedSpace.lean | 26 ++++---- .../LocallyRingedSpace/HasColimits.lean | 40 +++++++------ .../LocallyRingedSpace/ResidueField.lean | 11 ---- .../Geometry/RingedSpace/OpenImmersion.lean | 7 +-- .../LinearAlgebra/ExteriorAlgebra/Basic.lean | 8 ++- Mathlib/RingTheory/Henselian.lean | 9 ++- Mathlib/RingTheory/Ideal/Basic.lean | 2 +- .../LocalRing/ResidueField/Basic.lean | 27 +++++---- .../RingTheory/LocalRing/RingHom/Basic.lean | 47 ++++++++++----- Mathlib/RingTheory/Localization/AtPrime.lean | 10 +++- Mathlib/RingTheory/MvPowerSeries/Inverse.lean | 10 +++- Mathlib/RingTheory/PowerSeries/Inverse.lean | 10 +++- Mathlib/RingTheory/SurjectiveOnStalks.lean | 5 +- .../RingTheory/Valuation/ValExtension.lean | 10 +++- .../Valuation/ValuationSubring.lean | 2 +- 32 files changed, 269 insertions(+), 168 deletions(-) diff --git a/Mathlib/Algebra/Associated/Basic.lean b/Mathlib/Algebra/Associated/Basic.lean index a5cad90ce7c77..503074287a0cc 100644 --- a/Mathlib/Algebra/Associated/Basic.lean +++ b/Mathlib/Algebra/Associated/Basic.lean @@ -249,7 +249,7 @@ theorem Irreducible.dvd_comm [Monoid M] {p q : M} (hp : Irreducible p) (hq : Irr ⟨hp.dvd_symm hq, hq.dvd_symm hp⟩ theorem Irreducible.of_map {F : Type*} [Monoid M] [Monoid N] [FunLike F M N] [MonoidHomClass F M N] - {f : F} [IsLocalRingHom f] {x} (hfx : Irreducible (f x)) : Irreducible x := + {f : F} [IsLocalHom f] {x} (hfx : Irreducible (f x)) : Irreducible x := ⟨fun hu ↦ hfx.not_unit <| hu.map f, by rintro p q rfl exact (hfx.isUnit_or_isUnit <| map_mul f p q).imp (.of_map f _) (.of_map f _)⟩ diff --git a/Mathlib/Algebra/Category/Ring/Constructions.lean b/Mathlib/Algebra/Category/Ring/Constructions.lean index c2cc69f71a187..d9bb67a3267cd 100644 --- a/Mathlib/Algebra/Category/Ring/Constructions.lean +++ b/Mathlib/Algebra/Category/Ring/Constructions.lean @@ -224,7 +224,7 @@ def equalizerForkIsLimit : IsLimit (equalizerFork f g) := by · intro m hm exact RingHom.ext fun x => Subtype.ext <| ConcreteCategory.congr_hom hm x -instance : IsLocalRingHom (equalizerFork f g).ι := by +instance : IsLocalHom (equalizerFork f g).ι := by constructor rintro ⟨a, h₁ : _ = _⟩ (⟨⟨x, y, h₃, h₄⟩, rfl : x = _⟩ : IsUnit a) have : y ∈ RingHom.eqLocus f g := by @@ -234,8 +234,9 @@ instance : IsLocalRingHom (equalizerFork f g).ι := by rw [isUnit_iff_exists_inv] exact ⟨⟨y, this⟩, Subtype.eq h₃⟩ -instance equalizer_ι_isLocalRingHom (F : WalkingParallelPair ⥤ CommRingCat.{u}) : - IsLocalRingHom (limit.π F WalkingParallelPair.zero) := by +@[instance] +theorem equalizer_ι_isLocalHom (F : WalkingParallelPair ⥤ CommRingCat.{u}) : + IsLocalHom (limit.π F WalkingParallelPair.zero) := by have := limMap_π (diagramIsoParallelPair F).hom WalkingParallelPair.zero rw [← IsIso.comp_inv_eq] at this rw [← this] @@ -244,15 +245,18 @@ instance equalizer_ι_isLocalRingHom (F : WalkingParallelPair ⥤ CommRingCat.{u equalizerForkIsLimit (F.map WalkingParallelPairHom.left) (F.map WalkingParallelPairHom.right)⟩ WalkingParallelPair.zero] - change IsLocalRingHom ((lim.map _ ≫ _ ≫ (equalizerFork _ _).ι) ≫ _) + change IsLocalHom ((lim.map _ ≫ _ ≫ (equalizerFork _ _).ι) ≫ _) infer_instance +@[deprecated (since := "2024-10-10")] +alias equalizer_ι_isLocalRingHom := equalizer_ι_isLocalHom + open CategoryTheory.Limits.WalkingParallelPair Opposite open CategoryTheory.Limits.WalkingParallelPairHom instance equalizer_ι_is_local_ring_hom' (F : WalkingParallelPairᵒᵖ ⥤ CommRingCat.{u}) : - IsLocalRingHom (limit.π F (Opposite.op WalkingParallelPair.one)) := by + IsLocalHom (limit.π F (Opposite.op WalkingParallelPair.one)) := by have : _ = limit.π F (walkingParallelPairOpEquiv.functor.obj _) := (limit.isoLimitCone_inv_π ⟨_, IsLimit.whiskerEquivalence (limit.isLimit F) walkingParallelPairOpEquiv⟩ diff --git a/Mathlib/Algebra/Category/Ring/Instances.lean b/Mathlib/Algebra/Category/Ring/Instances.lean index fc588d27c8248..def5d4f5edf11 100644 --- a/Mathlib/Algebra/Category/Ring/Instances.lean +++ b/Mathlib/Algebra/Category/Ring/Instances.lean @@ -11,6 +11,7 @@ import Mathlib.RingTheory.LocalRing.RingHom.Basic # Ring-theoretic results in terms of categorical languages -/ +universe u open CategoryTheory @@ -36,16 +37,53 @@ instance Localization.epi' {R : CommRingCat} (M : Submonoid R) : rcases R with ⟨α, str⟩ exact IsLocalization.epi M _ -instance CommRingCat.isLocalRingHom_comp {R S T : CommRingCat} (f : R ⟶ S) (g : S ⟶ T) - [IsLocalRingHom g] [IsLocalRingHom f] : IsLocalRingHom (f ≫ g) := - RingHom.isLocalRingHom_comp _ _ +/- +These three instances solve the issue where the `FunLike` instances provided by +`CommRingCat.instFunLike'`, `CommRingCat.instFunLike''`, and `CommRingCat.instFunLike'''` +are not syntactically equal to `CommRingCat.instFunLike` when applied to +objects of the form `CommRingCat.of R`. +To prevent infinite loops, the priority of these three instances must be set lower +than that of other instances. +-/ +instance (priority := 50) {R : Type*} [CommRing R] {S : CommRingCat} (f : CommRingCat.of R ⟶ S) + [IsLocalHom (R := CommRingCat.of R) f] : IsLocalHom f := + inferInstance + +instance (priority := 50) {R : CommRingCat} {S : Type*} [CommRing S] (f : R ⟶ CommRingCat.of S) + [IsLocalHom (S := CommRingCat.of S) f] : IsLocalHom f := + inferInstance + +instance (priority := 50) {R S : Type u} [CommRing R] [CommRing S] + (f : CommRingCat.of R ⟶ CommRingCat.of S) + [IsLocalHom (R := CommRingCat.of R) (S := CommRingCat.of S) f] : IsLocalHom f := + inferInstance + +-- This instance handles the coercion of a morphism into a real `RingHom`. +instance {R S : CommRingCat} (f : R ⟶ S) [IsLocalHom f] : + IsLocalHom (F := R →+* S) f := + inferInstance -theorem isLocalRingHom_of_iso {R S : CommRingCat} (f : R ≅ S) : IsLocalRingHom f.hom := +@[instance] +theorem CommRingCat.isLocalHom_comp {R S T : CommRingCat} (f : R ⟶ S) (g : S ⟶ T) + [IsLocalHom g] [IsLocalHom f] : IsLocalHom (f ≫ g) := + RingHom.isLocalHom_comp _ _ + +@[deprecated (since := "2024-10-10")] +alias CommRingCat.isLocalRingHom_comp := CommRingCat.isLocalHom_comp + +theorem isLocalHom_of_iso {R S : CommRingCat} (f : R ≅ S) : IsLocalHom f.hom := { map_nonunit := fun a ha => by convert f.inv.isUnit_map ha exact (RingHom.congr_fun f.hom_inv_id _).symm } +@[deprecated (since := "2024-10-10")] +alias isLocalRingHom_of_iso := isLocalHom_of_iso + -- see Note [lower instance priority] -instance (priority := 100) isLocalRingHom_of_isIso {R S : CommRingCat} (f : R ⟶ S) [IsIso f] : - IsLocalRingHom f := - isLocalRingHom_of_iso (asIso f) +@[instance 100] +theorem isLocalHom_of_isIso {R S : CommRingCat} (f : R ⟶ S) [IsIso f] : + IsLocalHom f := + isLocalHom_of_iso (asIso f) + +@[deprecated (since := "2024-10-10")] +alias isLocalRingHom_of_isIso := isLocalHom_of_isIso diff --git a/Mathlib/Algebra/Group/Units/Equiv.lean b/Mathlib/Algebra/Group/Units/Equiv.lean index 12d6fb5b2285c..27f2a366dd06d 100644 --- a/Mathlib/Algebra/Group/Units/Equiv.lean +++ b/Mathlib/Algebra/Group/Units/Equiv.lean @@ -193,9 +193,13 @@ theorem MulEquiv.inv_symm (G : Type*) [DivisionCommMonoid G] : (MulEquiv.inv G).symm = MulEquiv.inv G := rfl -instance isLocalRingHom_equiv [Monoid M] [Monoid N] [EquivLike F M N] - [MulEquivClass F M N] (f : F) : IsLocalRingHom f where +@[instance] +theorem isLocalHom_equiv [Monoid M] [Monoid N] [EquivLike F M N] + [MulEquivClass F M N] (f : F) : IsLocalHom f where map_nonunit a ha := by convert ha.map (f : M ≃* N).symm rw [MulEquiv.eq_symm_apply] rfl -- note to reviewers: ugly `rfl` + +@[deprecated (since := "2024-10-10")] +alias isLocalRingHom_equiv := isLocalHom_equiv diff --git a/Mathlib/Algebra/Group/Units/Hom.lean b/Mathlib/Algebra/Group/Units/Hom.lean index 844235b25ef68..5a210213562a4 100644 --- a/Mathlib/Algebra/Group/Units/Hom.lean +++ b/Mathlib/Algebra/Group/Units/Hom.lean @@ -17,7 +17,7 @@ also contains unrelated results about `Units` that depend on `MonoidHom`. * `Units.map`: Turn a homomorphism from `α` to `β` monoids into a homomorphism from `αˣ` to `βˣ`. * `MonoidHom.toHomUnits`: Turn a homomorphism from a group `α` to `β` into a homomorphism from `α` to `βˣ`. -* `IsLocalRingHom`: A predicate on monoid maps, requiring that it maps nonunits +* `IsLocalHom`: A predicate on monoid maps, requiring that it maps nonunits to nonunits. For local rings, this means that the image of the unique maximal ideal is again contained in the unique maximal ideal. This is developed earlier, and in the generality of monoids, as it allows its use in non-local-ring related contexts, but it does have the @@ -28,7 +28,7 @@ also contains unrelated results about `Units` that depend on `MonoidHom`. The results that don't mention homomorphisms should be proved (earlier?) in a different file and be used to golf the basic `Group` lemmas. -Add a `@[to_additive]` version of `IsLocalRingHom`. +Add a `@[to_additive]` version of `IsLocalHom`. -/ assert_not_exists MonoidWithZero @@ -179,7 +179,7 @@ theorem of_leftInverse [MonoidHomClass G N M] {f : F} {x : M} (g : G) (hfg : Function.LeftInverse g f) (h : IsUnit (f x)) : IsUnit x := by simpa only [hfg x] using h.map g -/-- Prefer `IsLocalRingHom.of_leftInverse`, but we can't get rid of this because of `ToAdditive`. -/ +/-- Prefer `IsLocalHom.of_leftInverse`, but we can't get rid of this because of `ToAdditive`. -/ @[to_additive] theorem _root_.isUnit_map_of_leftInverse [MonoidHomClass F M N] [MonoidHomClass G N M] {f : F} {x : M} (g : G) (hfg : Function.LeftInverse g f) : @@ -208,7 +208,7 @@ theorem liftRight_inv_mul (f : M →* N) (h : ∀ x, IsUnit (f x)) (x) : end Monoid end IsUnit -section IsLocalRingHom +section IsLocalHom variable {G R S T F : Type*} @@ -218,13 +218,16 @@ variable [Monoid R] [Monoid S] [Monoid T] [FunLike F R S] is a unit if `f a` is a unit for any `a`. See `LocalRing.local_hom_TFAE` for other equivalent definitions in the local ring case - from where this concept originates, but it is useful in other contexts, so we allow this generalisation in mathlib. -/ -class IsLocalRingHom (f : F) : Prop where - /-- A local ring homomorphism `f : R ⟶ S` will send nonunits of `R` to nonunits of `S`. -/ +class IsLocalHom (f : F) : Prop where + /-- A local homomorphism `f : R ⟶ S` will send nonunits of `R` to nonunits of `S`. -/ map_nonunit : ∀ a, IsUnit (f a) → IsUnit a +@[deprecated (since := "2024-10-10")] +alias IsLocalRingHom := IsLocalHom + @[simp] -theorem IsUnit.of_map (f : F) [IsLocalRingHom f] (a : R) (h : IsUnit (f a)) : IsUnit a := - IsLocalRingHom.map_nonunit a h +theorem IsUnit.of_map (f : F) [IsLocalHom f] (a : R) (h : IsUnit (f a)) : IsUnit a := + IsLocalHom.map_nonunit a h -- TODO : remove alias, change the parenthesis of `f` and `a` alias isUnit_of_map_unit := IsUnit.of_map @@ -232,24 +235,38 @@ alias isUnit_of_map_unit := IsUnit.of_map variable [MonoidHomClass F R S] @[simp] -theorem isUnit_map_iff (f : F) [IsLocalRingHom f] (a : R) : IsUnit (f a) ↔ IsUnit a := - ⟨IsLocalRingHom.map_nonunit a, IsUnit.map f⟩ +theorem isUnit_map_iff (f : F) [IsLocalHom f] (a : R) : IsUnit (f a) ↔ IsUnit a := + ⟨IsLocalHom.map_nonunit a, IsUnit.map f⟩ -theorem isLocalRingHom_of_leftInverse [FunLike G S R] [MonoidHomClass G S R] - {f : F} (g : G) (hfg : Function.LeftInverse g f) : IsLocalRingHom f where +theorem isLocalHom_of_leftInverse [FunLike G S R] [MonoidHomClass G S R] + {f : F} (g : G) (hfg : Function.LeftInverse g f) : IsLocalHom f where map_nonunit a ha := by rwa [isUnit_map_of_leftInverse g hfg] at ha -instance MonoidHom.isLocalRingHom_comp (g : S →* T) (f : R →* S) [IsLocalRingHom g] - [IsLocalRingHom f] : IsLocalRingHom (g.comp f) where - map_nonunit a := IsLocalRingHom.map_nonunit a ∘ IsLocalRingHom.map_nonunit (f := g) (f a) +@[deprecated (since := "2024-10-10")] +alias isLocalRingHom_of_leftInverse := isLocalHom_of_leftInverse + +@[instance] +theorem MonoidHom.isLocalHom_comp (g : S →* T) (f : R →* S) [IsLocalHom g] + [IsLocalHom f] : IsLocalHom (g.comp f) where + map_nonunit a := IsLocalHom.map_nonunit a ∘ IsLocalHom.map_nonunit (f := g) (f a) + +@[deprecated (since := "2024-10-10")] +alias MonoidHom.isLocalRingHom_comp := MonoidHom.isLocalHom_comp -- see note [lower instance priority] -instance (priority := 100) isLocalRingHom_toMonoidHom (f : F) [IsLocalRingHom f] : - IsLocalRingHom (f : R →* S) := - ⟨IsLocalRingHom.map_nonunit (f := f)⟩ +@[instance 100] +theorem isLocalHom_toMonoidHom (f : F) [IsLocalHom f] : + IsLocalHom (f : R →* S) := + ⟨IsLocalHom.map_nonunit (f := f)⟩ -theorem MonoidHom.isLocalRingHom_of_comp (f : R →* S) (g : S →* T) [IsLocalRingHom (g.comp f)] : - IsLocalRingHom f := +@[deprecated (since := "2024-10-10")] +alias isLocalRingHom_toMonoidHom := isLocalHom_toMonoidHom + +theorem MonoidHom.isLocalHom_of_comp (f : R →* S) (g : S →* T) [IsLocalHom (g.comp f)] : + IsLocalHom f := ⟨fun _ ha => (isUnit_map_iff (g.comp f) _).mp (ha.map g)⟩ -end IsLocalRingHom +@[deprecated (since := "2024-10-10")] +alias MonoidHom.isLocalRingHom_of_comp := MonoidHom.isLocalHom_of_comp + +end IsLocalHom diff --git a/Mathlib/Algebra/GroupWithZero/Units/Lemmas.lean b/Mathlib/Algebra/GroupWithZero/Units/Lemmas.lean index 407a0fb1b6c75..2d5790637992e 100644 --- a/Mathlib/Algebra/GroupWithZero/Units/Lemmas.lean +++ b/Mathlib/Algebra/GroupWithZero/Units/Lemmas.lean @@ -22,8 +22,8 @@ section Monoid variable [Monoid M] [GroupWithZero G₀] -lemma isLocalRingHom_of_exists_map_ne_one [FunLike F G₀ M] [MonoidHomClass F G₀ M] {f : F} - (hf : ∃ x : G₀, f x ≠ 1) : IsLocalRingHom f where +lemma isLocalHom_of_exists_map_ne_one [FunLike F G₀ M] [MonoidHomClass F G₀ M] {f : F} + (hf : ∃ x : G₀, f x ≠ 1) : IsLocalHom f where map_nonunit a h := by rcases eq_or_ne a 0 with (rfl | h) · obtain ⟨t, ht⟩ := hf @@ -33,9 +33,12 @@ lemma isLocalRingHom_of_exists_map_ne_one [FunLike F G₀ M] [MonoidHomClass F G exact (h.mul_right_cancel this).symm · exact ⟨⟨a, a⁻¹, mul_inv_cancel₀ h, inv_mul_cancel₀ h⟩, rfl⟩ +@[deprecated (since := "2024-10-10")] +alias isLocalRingHom_of_exists_map_ne_one := isLocalHom_of_exists_map_ne_one + instance [GroupWithZero G₀] [FunLike F G₀ M₀] [MonoidWithZeroHomClass F G₀ M₀] [Nontrivial M₀] - (f : F) : IsLocalRingHom f := - isLocalRingHom_of_exists_map_ne_one ⟨0, by simp⟩ + (f : F) : IsLocalHom f := + isLocalHom_of_exists_map_ne_one ⟨0, by simp⟩ end Monoid diff --git a/Mathlib/Algebra/Polynomial/Expand.lean b/Mathlib/Algebra/Polynomial/Expand.lean index ae212d4a08496..33ed390954f87 100644 --- a/Mathlib/Algebra/Polynomial/Expand.lean +++ b/Mathlib/Algebra/Polynomial/Expand.lean @@ -268,17 +268,20 @@ section IsDomain variable (R : Type u) [CommRing R] [IsDomain R] -theorem isLocalRingHom_expand {p : ℕ} (hp : 0 < p) : IsLocalRingHom (expand R p) := by +theorem isLocalHom_expand {p : ℕ} (hp : 0 < p) : IsLocalHom (expand R p) := by refine ⟨fun f hf1 => ?_⟩ have hf2 := eq_C_of_degree_eq_zero (degree_eq_zero_of_isUnit hf1) rw [coeff_expand hp, if_pos (dvd_zero _), p.zero_div] at hf2 rw [hf2, isUnit_C] at hf1; rw [expand_eq_C hp] at hf2; rwa [hf2, isUnit_C] +@[deprecated (since := "2024-10-10")] +alias isLocalRingHom_expand := isLocalHom_expand + variable {R} theorem of_irreducible_expand {p : ℕ} (hp : p ≠ 0) {f : R[X]} (hf : Irreducible (expand R p f)) : Irreducible f := - let _ := isLocalRingHom_expand R hp.bot_lt + let _ := isLocalHom_expand R hp.bot_lt hf.of_map theorem of_irreducible_expand_pow {p : ℕ} (hp : p ≠ 0) {f : R[X]} {n : ℕ} : diff --git a/Mathlib/AlgebraicGeometry/AffineScheme.lean b/Mathlib/AlgebraicGeometry/AffineScheme.lean index 0859e4051faef..74ea6759b202f 100644 --- a/Mathlib/AlgebraicGeometry/AffineScheme.lean +++ b/Mathlib/AlgebraicGeometry/AffineScheme.lean @@ -288,7 +288,6 @@ lemma isoSpec_hom_base_apply (x : U) : (Iso.eq_comp_inv _).mpr (Scheme.Opens.germ_stalkIso_hom U (V := ⊤) x trivial), X.presheaf.germ_res_assoc, Spec.map_comp, Scheme.comp_base_apply] congr 1 - have := isLocalRingHom_of_isIso (U.stalkIso x).inv exact LocalRing.comap_closedPoint (U.stalkIso x).inv lemma isoSpec_inv_app_top : diff --git a/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean b/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean index bd775fd8d0f99..577af21ca59d1 100644 --- a/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean +++ b/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean @@ -271,8 +271,6 @@ def identityToΓSpec : 𝟭 LocallyRingedSpace.{u} ⟶ Γ.rightOp ⋙ Spec.toLoc dsimp show PrimeSpectrum.comap (f.c.app (op ⊤)) (X.toΓSpecFun x) = Y.toΓSpecFun (f.base x) dsimp [toΓSpecFun] - -- TODO: this instance was found automatically before #6045 - have := @AlgebraicGeometry.LocallyRingedSpace.isLocalRingHomStalkMap X Y rw [← LocalRing.comap_closedPoint (f.stalkMap x), ← PrimeSpectrum.comap_comp_apply, ← PrimeSpectrum.comap_comp_apply] congr 2 diff --git a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean index 26c8303a68816..05c884eb52c60 100644 --- a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean +++ b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean @@ -624,17 +624,20 @@ def closedPoint : PrimeSpectrum R := variable {R} -theorem isLocalRingHom_iff_comap_closedPoint {S : Type v} [CommSemiring S] [LocalRing S] - (f : R →+* S) : IsLocalRingHom f ↔ PrimeSpectrum.comap f (closedPoint S) = closedPoint R := by +theorem isLocalHom_iff_comap_closedPoint {S : Type v} [CommSemiring S] [LocalRing S] + (f : R →+* S) : IsLocalHom f ↔ PrimeSpectrum.comap f (closedPoint S) = closedPoint R := by -- Porting note: inline `this` does **not** work have := (local_hom_TFAE f).out 0 4 rw [this, PrimeSpectrum.ext_iff] rfl +@[deprecated (since := "2024-10-10")] +alias isLocalRingHom_iff_comap_closedPoint := isLocalHom_iff_comap_closedPoint + @[simp] theorem comap_closedPoint {S : Type v} [CommSemiring S] [LocalRing S] (f : R →+* S) - [IsLocalRingHom f] : PrimeSpectrum.comap f (closedPoint S) = closedPoint R := - (isLocalRingHom_iff_comap_closedPoint f).mp inferInstance + [IsLocalHom f] : PrimeSpectrum.comap f (closedPoint S) = closedPoint R := + (isLocalHom_iff_comap_closedPoint f).mp inferInstance theorem specializes_closedPoint (x : PrimeSpectrum R) : x ⤳ closedPoint R := (PrimeSpectrum.le_iff_specializes _ _).mp (LocalRing.le_maximalIdeal x.2.1) diff --git a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean index 106326628a9e4..77da59dbcba24 100644 --- a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean +++ b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean @@ -642,7 +642,7 @@ lemma toSpec_base_apply_eq_comap {f} (x : Proj| pbo f) : rw [awayToΓ_ΓToStalk, CommRingCat.comp_eq_ring_hom_comp, PrimeSpectrum.comap_comp] exact congr(PrimeSpectrum.comap _ $(@LocalRing.comap_closedPoint (HomogeneousLocalization.AtPrime 𝒜 x.1.asHomogeneousIdeal.toIdeal) _ _ - ((Proj| pbo f).presheaf.stalk x) _ _ _ (isLocalRingHom_of_isIso _))) + ((Proj| pbo f).presheaf.stalk x) _ _ _ (isLocalHom_of_isIso _))) lemma toSpec_base_apply_eq {f} (x : Proj| pbo f) : (toSpec 𝒜 f).base x = ProjIsoSpecTopComponent.toSpec 𝒜 f x := diff --git a/Mathlib/AlgebraicGeometry/ResidueField.lean b/Mathlib/AlgebraicGeometry/ResidueField.lean index a429de3ad3174..ac1d5d1e8dcc2 100644 --- a/Mathlib/AlgebraicGeometry/ResidueField.lean +++ b/Mathlib/AlgebraicGeometry/ResidueField.lean @@ -83,15 +83,6 @@ lemma basicOpen_eq_bot_iff_forall_evaluation_eq_zero (f : X.presheaf.obj (op U)) variable {X Y : Scheme.{u}} (f : X ⟶ Y) - --- TODO: This instance is found before #6045. --- We need this strange instance for `residueFieldMap`, the type of `F` must be fixed --- like this. The instance `IsLocalRingHom (f.stalkMap x)` already exists, but does not work for --- `residueFieldMap`. -instance (x): IsLocalRingHom (F := Y.presheaf.stalk (f.base x) →+* X.presheaf.stalk x) - (f.stalkMap x) := - f.1.2 x - /-- If `X ⟶ Y` is a morphism of locally ringed spaces and `x` a point of `X`, we obtain a morphism of residue fields in the other direction. -/ def Hom.residueFieldMap (f : X.Hom Y) (x : X) : diff --git a/Mathlib/AlgebraicGeometry/Scheme.lean b/Mathlib/AlgebraicGeometry/Scheme.lean index b958a149bf9ec..d77a37bef9383 100644 --- a/Mathlib/AlgebraicGeometry/Scheme.lean +++ b/Mathlib/AlgebraicGeometry/Scheme.lean @@ -640,6 +640,9 @@ namespace Scheme variable {X Y : Scheme.{u}} (f : X ⟶ Y) +instance (x) : IsLocalHom (f.stalkMap x) := + f.prop x + @[simp] lemma stalkMap_id (X : Scheme.{u}) (x : X) : (𝟙 X : X ⟶ X).stalkMap x = 𝟙 (X.presheaf.stalk x) := @@ -723,7 +726,7 @@ open LocalRing @[simp] lemma Spec_closedPoint {R S : CommRingCat} [LocalRing R] [LocalRing S] - {f : R ⟶ S} [IsLocalRingHom f] : (Spec.map f).base (closedPoint S) = closedPoint R := + {f : R ⟶ S} [IsLocalHom f] : (Spec.map f).base (closedPoint S) = closedPoint R := LocalRing.comap_closedPoint f end LocalRing diff --git a/Mathlib/AlgebraicGeometry/Spec.lean b/Mathlib/AlgebraicGeometry/Spec.lean index 25719eb7d693f..4d4595fade0bd 100644 --- a/Mathlib/AlgebraicGeometry/Spec.lean +++ b/Mathlib/AlgebraicGeometry/Spec.lean @@ -230,7 +230,7 @@ The induced map of a ring homomorphism on the prime spectra, as a morphism of lo def Spec.locallyRingedSpaceMap {R S : CommRingCat.{u}} (f : R ⟶ S) : Spec.locallyRingedSpaceObj S ⟶ Spec.locallyRingedSpaceObj R := LocallyRingedSpace.Hom.mk (Spec.sheafedSpaceMap f) fun p => - IsLocalRingHom.mk fun a ha => by + IsLocalHom.mk fun a ha => by -- Here, we are showing that the map on prime spectra induced by `f` is really a morphism of -- *locally* ringed spaces, i.e. that the induced map on the stalks is a local ring -- homomorphism. @@ -238,11 +238,9 @@ def Spec.locallyRingedSpaceMap {R S : CommRingCat.{u}} (f : R ⟶ S) : #adaptation_note /-- nightly-2024-04-01 It's this `erw` that is blowing up. The implicit arguments differ significantly. -/ erw [← localRingHom_comp_stalkIso_apply] at ha - -- TODO: this instance was found automatically before #6045 - haveI : IsLocalRingHom (stalkIso (↑S) p).inv := isLocalRingHom_of_isIso _ replace ha := (isUnit_map_iff (stalkIso S p).inv _).mp ha -- Porting note: `f` had to be made explicit - replace ha := IsLocalRingHom.map_nonunit + replace ha := IsLocalHom.map_nonunit (f := (Localization.localRingHom (PrimeSpectrum.comap f p).asIdeal p.asIdeal f _)) _ ha convert RingHom.isUnit_map (stalkIso R (PrimeSpectrum.comap f p)).inv ha erw [← comp_apply, show stalkToFiberRingHom R _ = (stalkIso _ _).hom from rfl, diff --git a/Mathlib/AlgebraicGeometry/Stalk.lean b/Mathlib/AlgebraicGeometry/Stalk.lean index c5d01c4906552..251d379d23a14 100644 --- a/Mathlib/AlgebraicGeometry/Stalk.lean +++ b/Mathlib/AlgebraicGeometry/Stalk.lean @@ -198,11 +198,9 @@ def stalkClosedPointTo : X.presheaf.stalk (f.base (closedPoint R)) ⟶ R := f.stalkMap (closedPoint R) ≫ (stalkClosedPointIso R).hom -instance isLocalRingHom_stalkClosedPointTo : - IsLocalRingHom (stalkClosedPointTo f) := by - apply (config := { allowSynthFailures := true }) RingHom.isLocalRingHom_comp - · apply isLocalRingHom_of_iso - · apply f.prop +instance isLocalHom_stalkClosedPointTo : + IsLocalHom (stalkClosedPointTo f) := + inferInstanceAs <| IsLocalHom (f.stalkMap (closedPoint R) ≫ (stalkClosedPointIso R).hom) lemma preimage_eq_top_of_closedPoint_mem {U : Opens X} (hU : f.base (closedPoint R) ∈ U) : f ⁻¹ᵁ U = ⊤ := @@ -234,7 +232,7 @@ lemma germ_stalkClosedPointTo (U : Opens X) (hU : f.base (closedPoint R) ∈ U) @[reassoc] lemma germ_stalkClosedPointTo_Spec_fromSpecStalk - {x : X} (f : X.presheaf.stalk x ⟶ R) [IsLocalRingHom f] (U : Opens X) (hU) : + {x : X} (f : X.presheaf.stalk x ⟶ R) [IsLocalHom f] (U : Opens X) (hU) : X.presheaf.germ U _ hU ≫ stalkClosedPointTo (Spec.map f ≫ X.fromSpecStalk x) = X.presheaf.germ U x (by simpa using hU) ≫ f := by have : (Spec.map f ≫ X.fromSpecStalk x).base (closedPoint R) = x := by @@ -278,7 +276,7 @@ variable {R} omit [LocalRing R] in /-- useful lemma for applications of `SpecToEquivOfLocalRing` -/ lemma SpecToEquivOfLocalRing_eq_iff - {f₁ f₂ : Σ x, { f : X.presheaf.stalk x ⟶ R // IsLocalRingHom f }} : + {f₁ f₂ : Σ x, { f : X.presheaf.stalk x ⟶ R // IsLocalHom f }} : f₁ = f₂ ↔ ∃ h₁ : f₁.1 = f₂.1, f₁.2.1 = (X.presheaf.stalkCongr (by rw [h₁]; rfl)).hom ≫ f₂.2.1 := by constructor @@ -297,7 +295,7 @@ Given a local ring `R` and scheme `X`, morphisms `Spec R ⟶ X` corresponds to p @[simps] noncomputable def SpecToEquivOfLocalRing : - (Spec R ⟶ X) ≃ Σ x, { f : X.presheaf.stalk x ⟶ R // IsLocalRingHom f } where + (Spec R ⟶ X) ≃ Σ x, { f : X.presheaf.stalk x ⟶ R // IsLocalHom f } where toFun f := ⟨f.base (closedPoint R), Scheme.stalkClosedPointTo f, inferInstance⟩ invFun xf := Spec.map xf.2.1 ≫ X.fromSpecStalk xf.1 left_inv := Scheme.Spec_stalkClosedPointTo_fromSpecStalk diff --git a/Mathlib/AlgebraicGeometry/StructureSheaf.lean b/Mathlib/AlgebraicGeometry/StructureSheaf.lean index 7541f38a34b7c..6d0abf3ae594a 100644 --- a/Mathlib/AlgebraicGeometry/StructureSheaf.lean +++ b/Mathlib/AlgebraicGeometry/StructureSheaf.lean @@ -542,14 +542,14 @@ def stalkIso (x : PrimeSpectrum.Top R) : instance (x : PrimeSpectrum R) : IsIso (stalkToFiberRingHom R x) := (stalkIso R x).isIso_hom -instance (x : PrimeSpectrum R) : IsLocalRingHom (stalkToFiberRingHom R x) := - isLocalRingHom_of_isIso _ +instance (x : PrimeSpectrum R) : IsLocalHom (stalkToFiberRingHom R x) := + isLocalHom_of_isIso _ instance (x : PrimeSpectrum R) : IsIso (localizationToStalk R x) := (stalkIso R x).isIso_inv -instance (x : PrimeSpectrum R) : IsLocalRingHom (localizationToStalk R x) := - isLocalRingHom_of_isIso _ +instance (x : PrimeSpectrum R) : IsLocalHom (localizationToStalk R x) := + isLocalHom_of_isIso _ @[simp, reassoc] theorem stalkToFiberRingHom_localizationToStalk (x : PrimeSpectrum.Top R) : diff --git a/Mathlib/FieldTheory/Separable.lean b/Mathlib/FieldTheory/Separable.lean index 63d3ad68e42f2..b8cfde792d439 100644 --- a/Mathlib/FieldTheory/Separable.lean +++ b/Mathlib/FieldTheory/Separable.lean @@ -350,7 +350,7 @@ theorem separable_or {f : F[X]} (hf : Irreducible f) : have := natDegree_eq_zero_of_derivative_eq_zero H have := (natDegree_pos_iff_degree_pos.mpr <| degree_pos_of_irreducible hf).ne' contradiction - haveI := isLocalRingHom_expand F hp + haveI := isLocalHom_expand F hp exact Or.inr ⟨by rw [separable_iff_derivative_ne_zero hf, Classical.not_not, H], contract p f, diff --git a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace.lean b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace.lean index 72d5eb095aa96..f0aedb017be00 100644 --- a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace.lean +++ b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace.lean @@ -11,7 +11,7 @@ import Mathlib.Geometry.RingedSpace.Stalks We define (bundled) locally ringed spaces (as `SheafedSpace CommRing` along with the fact that the stalks are local rings), and morphisms between these (morphisms in `SheafedSpace` with -`IsLocalRingHom` on the stalk maps). +`IsLocalHom` on the stalk maps). -/ -- Explicit universe annotations were used in this file to improve performance #12737 @@ -73,7 +73,7 @@ def 𝒪 : Sheaf CommRingCat X.toTopCat := structure Hom (X Y : LocallyRingedSpace.{u}) extends X.toPresheafedSpace.Hom Y.toPresheafedSpace : Type _ where /-- the underlying morphism induces a local ring homomorphism on stalks -/ - prop : ∀ x, IsLocalRingHom (toHom.stalkMap x) + prop : ∀ x, IsLocalHom (toHom.stalkMap x) /-- A morphism of locally ringed spaces as a morphism of sheafed spaces. -/ abbrev Hom.toShHom {X Y : LocallyRingedSpace.{u}} (f : X.Hom Y) : @@ -104,18 +104,23 @@ noncomputable def Hom.stalkMap {X Y : LocallyRingedSpace.{u}} (f : Hom X Y) (x : Y.presheaf.stalk (f.1.1 x) ⟶ X.presheaf.stalk x := f.toShHom.stalkMap x -instance isLocalRingHomStalkMap {X Y : LocallyRingedSpace.{u}} (f : X ⟶ Y) (x : X) : - IsLocalRingHom (f.stalkMap x) := +@[instance] +theorem isLocalHomStalkMap {X Y : LocallyRingedSpace.{u}} (f : X ⟶ Y) (x : X) : + IsLocalHom (f.stalkMap x) := f.2 x -instance isLocalRingHomValStalkMap {X Y : LocallyRingedSpace.{u}} (f : X ⟶ Y) (x : X) : - IsLocalRingHom (f.toShHom.stalkMap x) := +@[instance] +theorem isLocalHomValStalkMap {X Y : LocallyRingedSpace.{u}} (f : X ⟶ Y) (x : X) : + IsLocalHom (f.toShHom.stalkMap x) := f.2 x +@[deprecated (since := "2024-10-10")] +alias isLocalRingHomValStalkMap := isLocalHomValStalkMap + /-- The identity morphism on a locally ringed space. -/ @[simps! toShHom] def id (X : LocallyRingedSpace.{u}) : Hom X X := - ⟨𝟙 X.toSheafedSpace, fun x => by erw [PresheafedSpace.stalkMap.id]; apply isLocalRingHom_id⟩ + ⟨𝟙 X.toSheafedSpace, fun x => by erw [PresheafedSpace.stalkMap.id]; infer_instance⟩ instance (X : LocallyRingedSpace.{u}) : Inhabited (Hom X X) := ⟨id X⟩ @@ -124,7 +129,7 @@ instance (X : LocallyRingedSpace.{u}) : Inhabited (Hom X X) := def comp {X Y Z : LocallyRingedSpace.{u}} (f : Hom X Y) (g : Hom Y Z) : Hom X Z := ⟨f.toShHom ≫ g.toShHom, fun x => by erw [PresheafedSpace.stalkMap.comp] - exact @RingHom.isLocalRingHom_comp _ _ _ _ _ _ _ _ (f.2 _) (g.2 _)⟩ + infer_instance⟩ /-- The category of locally ringed spaces. -/ instance : Category LocallyRingedSpace.{u} where @@ -180,12 +185,11 @@ See also `isoOfSheafedSpaceIso`. @[simps! toShHom] def homOfSheafedSpaceHomOfIsIso {X Y : LocallyRingedSpace.{u}} (f : X.toSheafedSpace ⟶ Y.toSheafedSpace) [IsIso f] : X ⟶ Y := - Hom.mk f fun x => + Hom.mk f fun _ => -- Here we need to see that the stalk maps are really local ring homomorphisms. -- This can be solved by type class inference, because stalk maps of isomorphisms -- are isomorphisms and isomorphisms are local ring homomorphisms. - show IsLocalRingHom ((SheafedSpace.forgetToPresheafedSpace.map f).stalkMap x) by - infer_instance + inferInstance /-- Given two locally ringed spaces `X` and `Y`, an isomorphism between `X` and `Y` as _sheafed_ spaces can be lifted to an isomorphism `X ⟶ Y` as locally ringed spaces. diff --git a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean index a45f070a1bf4f..01ff44c393aaa 100644 --- a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean +++ b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean @@ -102,7 +102,7 @@ noncomputable def coproductCofanIsColimit : IsColimit (coproductCofan F) where (colimit.ι_desc (C := SheafedSpace.{u+1, u, u} CommRingCatMax.{u, u}) (forgetToSheafedSpace.mapCocone s) i : _)] haveI : - IsLocalRingHom + IsLocalHom (((forgetToSheafedSpace.mapCocone s).ι.app i).stalkMap y) := (s.ι.app i).2 y infer_instance⟩ @@ -131,9 +131,10 @@ variable {X Y : LocallyRingedSpace.{v}} (f g : X ⟶ Y) namespace HasCoequalizer -instance coequalizer_π_app_isLocalRingHom +@[instance] +theorem coequalizer_π_app_isLocalHom (U : TopologicalSpace.Opens (coequalizer f.toShHom g.toShHom).carrier) : - IsLocalRingHom ((coequalizer.π f.toShHom g.toShHom : _).c.app (op U)) := by + IsLocalHom ((coequalizer.π f.toShHom g.toShHom : _).c.app (op U)) := by have := ι_comp_coequalizerComparison f.toShHom g.toShHom SheafedSpace.forgetToPresheafedSpace rw [← PreservesCoequalizer.iso_hom] at this erw [SheafedSpace.congr_app this.symm (op U)] @@ -144,6 +145,9 @@ instance coequalizer_π_app_isLocalRingHom PresheafedSpace.c_isIso_of_iso _ infer_instance +@[deprecated (since := "2024-10-10")] +alias coequalizer_π_app_isLocalRingHom := coequalizer_π_app_isLocalHom + /-! We roughly follow the construction given in [MR0302656]. Given a pair `f, g : X ⟶ Y` of morphisms of locally ringed spaces, we want to show that the stalk map of @@ -210,8 +214,9 @@ theorem imageBasicOpen_image_open : erw [imageBasicOpen_image_preimage] exact (imageBasicOpen f g U s).2 -instance coequalizer_π_stalk_isLocalRingHom (x : Y) : - IsLocalRingHom ((coequalizer.π f.toShHom g.toShHom : _).stalkMap x) := by +@[instance] +theorem coequalizer_π_stalk_isLocalHom (x : Y) : + IsLocalHom ((coequalizer.π f.toShHom g.toShHom : _).stalkMap x) := by constructor rintro a ha rcases TopCat.Presheaf.germ_exist _ _ a with ⟨U, hU, s, rfl⟩ @@ -240,6 +245,9 @@ instance coequalizer_π_stalk_isLocalRingHom (x : Y) : convert @RingedSpace.isUnit_res_basicOpen Y.toRingedSpace (unop _) (((coequalizer.π f.toShHom g.toShHom).c.app (op U)) s) +@[deprecated (since := "2024-10-10")] +alias coequalizer_π_stalk_isLocalRingHom := coequalizer_π_stalk_isLocalHom + end HasCoequalizer /-- The coequalizer of two locally ringed space in the category of sheafed spaces is a locally @@ -249,20 +257,18 @@ noncomputable def coequalizer : LocallyRingedSpace where localRing x := by obtain ⟨y, rfl⟩ := (TopCat.epi_iff_surjective (coequalizer.π f.toShHom g.toShHom).base).mp inferInstance x - -- TODO: this instance was found automatically before #6045 - have _ : IsLocalRingHom ((coequalizer.π f.toShHom g.toShHom).stalkMap y) := inferInstance exact ((coequalizer.π f.toShHom g.toShHom : _).stalkMap y).domain_localRing /-- The explicit coequalizer cofork of locally ringed spaces. -/ noncomputable def coequalizerCofork : Cofork f g := @Cofork.ofπ _ _ _ _ f g (coequalizer f g) ⟨coequalizer.π f.toShHom g.toShHom, -- Porting note: this used to be automatic - HasCoequalizer.coequalizer_π_stalk_isLocalRingHom _ _⟩ + HasCoequalizer.coequalizer_π_stalk_isLocalHom _ _⟩ (LocallyRingedSpace.Hom.ext' (coequalizer.condition f.toShHom g.toShHom)) -theorem isLocalRingHom_stalkMap_congr {X Y : RingedSpace} (f g : X ⟶ Y) (H : f = g) (x) - (h : IsLocalRingHom (f.stalkMap x)) : - IsLocalRingHom (g.stalkMap x) := by +theorem isLocalHom_stalkMap_congr {X Y : RingedSpace} (f g : X ⟶ Y) (H : f = g) (x) + (h : IsLocalHom (f.stalkMap x)) : + IsLocalHom (g.stalkMap x) := by rw [PresheafedSpace.stalkMap.congr_hom _ _ H.symm x]; infer_instance /-- The cofork constructed in `coequalizer_cofork` is indeed a colimit cocone. -/ @@ -274,19 +280,19 @@ noncomputable def coequalizerCoforkIsColimit : IsColimit (coequalizerCofork f g) · intro x rcases (TopCat.epi_iff_surjective (coequalizer.π f.toShHom g.toShHom).base).mp inferInstance x with ⟨y, rfl⟩ - -- Porting note: was `apply isLocalRingHom_of_comp _ (PresheafedSpace.stalkMap ...)`, this + -- Porting note: was `apply isLocalHom_of_comp _ (PresheafedSpace.stalkMap ...)`, this -- used to allow you to provide the proof that `... ≫ ...` is a local ring homomorphism later, -- but this is no longer possible set h := _ - change IsLocalRingHom h - suffices _ : IsLocalRingHom (((coequalizerCofork f g).π.1.stalkMap _).comp h) by - apply isLocalRingHom_of_comp _ ((coequalizerCofork f g).π.1.stalkMap _) + change IsLocalHom h + suffices _ : IsLocalHom (((coequalizerCofork f g).π.1.stalkMap _).comp h) by + apply isLocalHom_of_comp _ ((coequalizerCofork f g).π.1.stalkMap _) -- note to reviewers: this `change` is now more brittle because it now has to fully resolve -- the type to be able to search for `MonoidHomClass`, even though of course all homs in -- `CommRingCat` are clearly such - change IsLocalRingHom (h ≫ (coequalizerCofork f g).π.toShHom.stalkMap y) + change IsLocalHom (h ≫ (coequalizerCofork f g).π.toShHom.stalkMap y) erw [← PresheafedSpace.stalkMap.comp] - apply isLocalRingHom_stalkMap_congr _ _ (coequalizer.π_desc s.π.toShHom e).symm y + apply isLocalHom_stalkMap_congr _ _ (coequalizer.π_desc s.π.toShHom e).symm y infer_instance constructor · exact LocallyRingedSpace.Hom.ext' (coequalizer.π_desc _ _) diff --git a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/ResidueField.lean b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/ResidueField.lean index 89243a71df33e..3a60a5f355de1 100644 --- a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/ResidueField.lean +++ b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/ResidueField.lean @@ -85,14 +85,6 @@ lemma Γevaluation_ne_zero_iff_mem_basicOpen (x : X) (f : X.presheaf.obj (op ⊤ variable {X Y : LocallyRingedSpace.{u}} (f : X ⟶ Y) (x : X) --- TODO: This instance is found before #6045. --- We need this strange instance for `residueFieldMap`, the type of `F` must be fixed --- like this. The instance `IsLocalRingHom (f.stalkMap x)` already exists, but does not work for --- `residueFieldMap`. -instance : IsLocalRingHom (F := Y.presheaf.stalk (f.base x) →+* X.presheaf.stalk x) - (f.stalkMap x) := - f.2 x - /-- If `X ⟶ Y` is a morphism of locally ringed spaces and `x` a point of `X`, we obtain a morphism of residue fields in the other direction. -/ def residueFieldMap (x : X) : Y.residueField (f.base x) ⟶ X.residueField x := @@ -114,9 +106,6 @@ lemma residueFieldMap_comp {Z : LocallyRingedSpace.{u}} (g : Y ⟶ Z) (x : X) : residueFieldMap (f ≫ g) x = residueFieldMap g (f.base x) ≫ residueFieldMap f x := by simp only [comp_toShHom, SheafedSpace.comp_base, Function.comp_apply, residueFieldMap] simp_rw [stalkMap_comp] - haveI : IsLocalRingHom (g.stalkMap (f.base x)) := inferInstance - -- TODO: This instance is found before #6045. - haveI : IsLocalRingHom (f.stalkMap x) := inferInstance apply LocalRing.ResidueField.map_comp @[reassoc] diff --git a/Mathlib/Geometry/RingedSpace/OpenImmersion.lean b/Mathlib/Geometry/RingedSpace/OpenImmersion.lean index a29fd4205a61a..959214c257295 100644 --- a/Mathlib/Geometry/RingedSpace/OpenImmersion.lean +++ b/Mathlib/Geometry/RingedSpace/OpenImmersion.lean @@ -947,7 +947,7 @@ instance : SheafedSpace.IsOpenImmersion (LocallyRingedSpace.forgetToSheafedSpace H -- note to reviewers: is there a `count_heartbeats` for this? -set_option synthInstance.maxHeartbeats 30000 in +set_option synthInstance.maxHeartbeats 40000 in /-- An explicit pullback cone over `cospan f g` if `f` is an open immersion. -/ def pullbackConeOfLeft : PullbackCone f g := by refine PullbackCone.mk ?_ @@ -966,7 +966,7 @@ def pullbackConeOfLeft : PullbackCone f g := by instance : LocallyRingedSpace.IsOpenImmersion (pullbackConeOfLeft f g).snd := show PresheafedSpace.IsOpenImmersion (Y.toPresheafedSpace.ofRestrict _) by infer_instance -set_option synthInstance.maxHeartbeats 80000 in +set_option synthInstance.maxHeartbeats 40000 in /-- The constructed `pullbackConeOfLeft` is indeed limiting. -/ def pullbackConeOfLeftIsLimit : IsLimit (pullbackConeOfLeft f g) := PullbackCone.isLimitAux' _ fun s => by @@ -986,8 +986,7 @@ def pullbackConeOfLeftIsLimit : IsLimit (pullbackConeOfLeft f g) := change _ = _ ≫ s.snd.1.stalkMap x at this rw [PresheafedSpace.stalkMap.comp, ← IsIso.eq_inv_comp] at this rw [this] - -- TODO: This instance is found by `infer_instance` before #6045. - apply CommRingCat.isLocalRingHom_comp + infer_instance · intro m _ h₂ rw [← cancel_mono (pullbackConeOfLeft f g).snd] exact h₂.trans <| LocallyRingedSpace.Hom.ext' diff --git a/Mathlib/LinearAlgebra/ExteriorAlgebra/Basic.lean b/Mathlib/LinearAlgebra/ExteriorAlgebra/Basic.lean index cd48d0d7611d4..258bd811ae3b5 100644 --- a/Mathlib/LinearAlgebra/ExteriorAlgebra/Basic.lean +++ b/Mathlib/LinearAlgebra/ExteriorAlgebra/Basic.lean @@ -164,8 +164,12 @@ theorem algebraMap_eq_zero_iff (x : R) : algebraMap R (ExteriorAlgebra R M) x = theorem algebraMap_eq_one_iff (x : R) : algebraMap R (ExteriorAlgebra R M) x = 1 ↔ x = 1 := map_eq_one_iff (algebraMap _ _) (algebraMap_leftInverse _).injective -instance isLocalRingHom_algebraMap : IsLocalRingHom (algebraMap R (ExteriorAlgebra R M)) := - isLocalRingHom_of_leftInverse _ (algebraMap_leftInverse M) +@[instance] +theorem isLocalHom_algebraMap : IsLocalHom (algebraMap R (ExteriorAlgebra R M)) := + isLocalHom_of_leftInverse _ (algebraMap_leftInverse M) + +@[deprecated (since := "2024-10-10")] +alias isLocalRingHom_algebraMap := isLocalHom_algebraMap theorem isUnit_algebraMap (r : R) : IsUnit (algebraMap R (ExteriorAlgebra R M) r) ↔ IsUnit r := isUnit_map_of_leftInverse _ (algebraMap_leftInverse M) diff --git a/Mathlib/RingTheory/Henselian.lean b/Mathlib/RingTheory/Henselian.lean index b1762c379f824..42918161d3de0 100644 --- a/Mathlib/RingTheory/Henselian.lean +++ b/Mathlib/RingTheory/Henselian.lean @@ -60,8 +60,8 @@ universe u v open Polynomial LocalRing Polynomial Function List -theorem isLocalRingHom_of_le_jacobson_bot {R : Type*} [CommRing R] (I : Ideal R) - (h : I ≤ Ideal.jacobson ⊥) : IsLocalRingHom (Ideal.Quotient.mk I) := by +theorem isLocalHom_of_le_jacobson_bot {R : Type*} [CommRing R] (I : Ideal R) + (h : I ≤ Ideal.jacobson ⊥) : IsLocalHom (Ideal.Quotient.mk I) := by constructor intro a h have : IsUnit (Ideal.Quotient.mk (Ideal.jacobson ⊥) a) := by @@ -79,6 +79,9 @@ theorem isLocalRingHom_of_le_jacobson_bot {R : Type*} [CommRing R] (I : Ideal R) simp? at h1 says simp only [mul_one, sub_add_cancel, IsUnit.mul_iff] at h1 exact h1.1 +@[deprecated (since := "2024-10-10")] +alias isLocalRingHom_of_le_jacobson_bot := isLocalHom_of_le_jacobson_bot + /-- A ring `R` is *Henselian* at an ideal `I` if the following condition holds: for every polynomial `f` over `R`, with a *simple* root `a₀` over the quotient ring `R/I`, there exists a lift `a : R` of `a₀` that is a root of `f`. @@ -193,7 +196,7 @@ instance (priority := 100) IsAdicComplete.henselianRing (R : Type*) [CommRing R] exact (ih.eval f).trans h₁ have hf'c : ∀ n, IsUnit (f'.eval (c n)) := by intro n - haveI := isLocalRingHom_of_le_jacobson_bot I (IsAdicComplete.le_jacobson_bot I) + haveI := isLocalHom_of_le_jacobson_bot I (IsAdicComplete.le_jacobson_bot I) apply IsUnit.of_map (Ideal.Quotient.mk I) convert h₂ using 1 exact SModEq.def.mp ((hc_mod n).eval _) diff --git a/Mathlib/RingTheory/Ideal/Basic.lean b/Mathlib/RingTheory/Ideal/Basic.lean index 69a42212b475e..1db54f75b9e6e 100644 --- a/Mathlib/RingTheory/Ideal/Basic.lean +++ b/Mathlib/RingTheory/Ideal/Basic.lean @@ -808,7 +808,7 @@ theorem one_not_mem_nonunits [Monoid α] : (1 : α) ∉ nonunits α := -- Porting note : as this can be proved by other `simp` lemmas, this is marked as high priority. @[simp (high)] theorem map_mem_nonunits_iff [Monoid α] [Monoid β] [FunLike F α β] [MonoidHomClass F α β] (f : F) - [IsLocalRingHom f] (a) : f a ∈ nonunits β ↔ a ∈ nonunits α := + [IsLocalHom f] (a) : f a ∈ nonunits β ↔ a ∈ nonunits α := ⟨fun h ha => h <| ha.map f, fun h ha => h <| ha.of_map⟩ theorem coe_subset_nonunits [Semiring α] {I : Ideal α} (h : I ≠ ⊤) : (I : Set α) ⊆ nonunits α := diff --git a/Mathlib/RingTheory/LocalRing/ResidueField/Basic.lean b/Mathlib/RingTheory/LocalRing/ResidueField/Basic.lean index ca82438849309..5dc19df864940 100644 --- a/Mathlib/RingTheory/LocalRing/ResidueField/Basic.lean +++ b/Mathlib/RingTheory/LocalRing/ResidueField/Basic.lean @@ -41,7 +41,7 @@ instance ResidueField.algebra : Algebra R (ResidueField R) := theorem ResidueField.algebraMap_eq : algebraMap R (ResidueField R) = residue R := rfl -instance : IsLocalRingHom (LocalRing.residue R) := +instance : IsLocalHom (LocalRing.residue R) := ⟨fun _ ha => Classical.not_not.mp (Ideal.Quotient.eq_zero_iff_mem.not.mp (isUnit_iff_ne_zero.mp ha))⟩ @@ -50,22 +50,22 @@ variable {R} namespace ResidueField /-- A local ring homomorphism into a field can be descended onto the residue field. -/ -def lift {R S : Type*} [CommRing R] [LocalRing R] [Field S] (f : R →+* S) [IsLocalRingHom f] : +def lift {R S : Type*} [CommRing R] [LocalRing R] [Field S] (f : R →+* S) [IsLocalHom f] : LocalRing.ResidueField R →+* S := Ideal.Quotient.lift _ f fun a ha => by_contradiction fun h => ha (isUnit_of_map_unit f a (isUnit_iff_ne_zero.mpr h)) theorem lift_comp_residue {R S : Type*} [CommRing R] [LocalRing R] [Field S] (f : R →+* S) - [IsLocalRingHom f] : (lift f).comp (residue R) = f := + [IsLocalHom f] : (lift f).comp (residue R) = f := RingHom.ext fun _ => rfl @[simp] theorem lift_residue_apply {R S : Type*} [CommRing R] [LocalRing R] [Field S] (f : R →+* S) - [IsLocalRingHom f] (x) : lift f (residue R x) = f x := + [IsLocalHom f] (x) : lift f (residue R x) = f x := rfl /-- The map on residue fields induced by a local homomorphism between local rings -/ -def map (f : R →+* S) [IsLocalRingHom f] : ResidueField R →+* ResidueField S := +def map (f : R →+* S) [IsLocalHom f] : ResidueField R →+* ResidueField S := Ideal.Quotient.lift (maximalIdeal R) ((Ideal.Quotient.mk _).comp f) fun a ha => by erw [Ideal.Quotient.eq_zero_iff_mem] exact map_nonunit f a ha @@ -79,16 +79,16 @@ theorem map_id : /-- The composite of two `LocalRing.ResidueField.map`s is the `LocalRing.ResidueField.map` of the composite. -/ -theorem map_comp (f : T →+* R) (g : R →+* S) [IsLocalRingHom f] [IsLocalRingHom g] : +theorem map_comp (f : T →+* R) (g : R →+* S) [IsLocalHom f] [IsLocalHom g] : LocalRing.ResidueField.map (g.comp f) = (LocalRing.ResidueField.map g).comp (LocalRing.ResidueField.map f) := Ideal.Quotient.ringHom_ext <| RingHom.ext fun _ => rfl -theorem map_comp_residue (f : R →+* S) [IsLocalRingHom f] : +theorem map_comp_residue (f : R →+* S) [IsLocalHom f] : (ResidueField.map f).comp (residue R) = (residue S).comp f := rfl -theorem map_residue (f : R →+* S) [IsLocalRingHom f] (r : R) : +theorem map_residue (f : R →+* S) [IsLocalHom f] (r : R) : ResidueField.map f (residue R r) = residue S (f r) := rfl @@ -96,8 +96,8 @@ theorem map_id_apply (x : ResidueField R) : map (RingHom.id R) x = x := DFunLike.congr_fun map_id x @[simp] -theorem map_map (f : R →+* S) (g : S →+* T) (x : ResidueField R) [IsLocalRingHom f] - [IsLocalRingHom g] : map g (map f x) = map (g.comp f) x := +theorem map_map (f : R →+* S) (g : S →+* T) (x : ResidueField R) [IsLocalHom f] + [IsLocalHom g] : map g (map f x) = map (g.comp f) x := DFunLike.congr_fun (map_comp f g).symm x /-- A ring isomorphism defines an isomorphism of residue fields. -/ @@ -147,7 +147,7 @@ end MulSemiringAction section FiniteDimensional -variable [Algebra R S] [IsLocalRingHom (algebraMap R S)] +variable [Algebra R S] [IsLocalHom (algebraMap R S)] noncomputable instance : Algebra (ResidueField R) (ResidueField S) := (ResidueField.map (algebraMap R S)).toAlgebra @@ -175,13 +175,16 @@ end FiniteDimensional end ResidueField -theorem isLocalRingHom_residue : IsLocalRingHom (LocalRing.residue R) := by +theorem isLocalHom_residue : IsLocalHom (LocalRing.residue R) := by constructor intro a ha by_contra h erw [Ideal.Quotient.eq_zero_iff_mem.mpr ((LocalRing.mem_maximalIdeal _).mpr h)] at ha exact ha.ne_zero rfl +@[deprecated (since := "2024-10-10")] +alias isLocalRingHom_residue := isLocalHom_residue + end end LocalRing diff --git a/Mathlib/RingTheory/LocalRing/RingHom/Basic.lean b/Mathlib/RingTheory/LocalRing/RingHom/Basic.lean index 4b44e62f3ee66..f14d328b4582a 100644 --- a/Mathlib/RingTheory/LocalRing/RingHom/Basic.lean +++ b/Mathlib/RingTheory/LocalRing/RingHom/Basic.lean @@ -21,25 +21,40 @@ section variable [Semiring R] [Semiring S] [Semiring T] -instance isLocalRingHom_id (R : Type*) [Semiring R] : IsLocalRingHom (RingHom.id R) where +@[instance] +theorem isLocalHom_id (R : Type*) [Semiring R] : IsLocalHom (RingHom.id R) where map_nonunit _ := id +@[deprecated (since := "2024-10-10")] +alias isLocalRingHom_id := isLocalHom_id + -- see note [lower instance priority] -instance (priority := 100) isLocalRingHom_toRingHom {F : Type*} [FunLike F R S] - [RingHomClass F R S] (f : F) [IsLocalRingHom f] : IsLocalRingHom (f : R →+* S) := - ⟨IsLocalRingHom.map_nonunit (f := f)⟩ +@[instance 100] +theorem isLocalHom_toRingHom {F : Type*} [FunLike F R S] + [RingHomClass F R S] (f : F) [IsLocalHom f] : IsLocalHom (f : R →+* S) := + ⟨IsLocalHom.map_nonunit (f := f)⟩ + +@[deprecated (since := "2024-10-10")] +alias isLocalRingHom_toRingHom := isLocalHom_toRingHom -instance RingHom.isLocalRingHom_comp (g : S →+* T) (f : R →+* S) [IsLocalRingHom g] - [IsLocalRingHom f] : IsLocalRingHom (g.comp f) where - map_nonunit a := IsLocalRingHom.map_nonunit a ∘ IsLocalRingHom.map_nonunit (f := g) (f a) +@[instance] +theorem RingHom.isLocalHom_comp (g : S →+* T) (f : R →+* S) [IsLocalHom g] + [IsLocalHom f] : IsLocalHom (g.comp f) where + map_nonunit a := IsLocalHom.map_nonunit a ∘ IsLocalHom.map_nonunit (f := g) (f a) -theorem isLocalRingHom_of_comp (f : R →+* S) (g : S →+* T) [IsLocalRingHom (g.comp f)] : - IsLocalRingHom f := +@[deprecated (since := "2024-10-10")] +alias RingHom.isLocalRingHom_comp := RingHom.isLocalHom_comp + +theorem isLocalHom_of_comp (f : R →+* S) (g : S →+* T) [IsLocalHom (g.comp f)] : + IsLocalHom f := ⟨fun _ ha => (isUnit_map_iff (g.comp f) _).mp (g.isUnit_map ha)⟩ +@[deprecated (since := "2024-10-10")] +alias isLocalRingHom_of_comp := isLocalHom_of_comp + /-- If `f : R →+* S` is a local ring hom, then `R` is a local ring if `S` is. -/ theorem RingHom.domain_localRing {R S : Type*} [CommSemiring R] [CommSemiring S] [H : LocalRing S] - (f : R →+* S) [IsLocalRingHom f] : LocalRing R := by + (f : R →+* S) [IsLocalHom f] : LocalRing R := by haveI : Nontrivial R := pullback_nonzero f f.map_zero f.map_one apply LocalRing.of_nonunits_add intro a b @@ -57,7 +72,7 @@ variable [CommSemiring R] [LocalRing R] [CommSemiring S] [LocalRing S] /-- The image of the maximal ideal of the source is contained within the maximal ideal of the target. -/ -theorem map_nonunit (f : R →+* S) [IsLocalRingHom f] (a : R) (h : a ∈ maximalIdeal R) : +theorem map_nonunit (f : R →+* S) [IsLocalHom f] (a : R) (h : a ∈ maximalIdeal R) : f a ∈ maximalIdeal S := fun H => h <| isUnit_of_map_unit f a H end @@ -73,7 +88,7 @@ i.e. any preimage of a unit is still a unit. https://stacks.math.columbia.edu/ta -/ theorem local_hom_TFAE (f : R →+* S) : List.TFAE - [IsLocalRingHom f, f '' (maximalIdeal R).1 ⊆ maximalIdeal S, + [IsLocalHom f, f '' (maximalIdeal R).1 ⊆ maximalIdeal S, (maximalIdeal R).map f ≤ maximalIdeal S, maximalIdeal R ≤ (maximalIdeal S).comap f, (maximalIdeal S).comap f = maximalIdeal R] := by tfae_have 1 → 2 @@ -89,19 +104,19 @@ theorem local_hom_TFAE (f : R →+* S) : end theorem of_surjective [CommSemiring R] [LocalRing R] [CommSemiring S] [Nontrivial S] (f : R →+* S) - [IsLocalRingHom f] (hf : Function.Surjective f) : LocalRing S := + [IsLocalHom f] (hf : Function.Surjective f) : LocalRing S := of_isUnit_or_isUnit_of_isUnit_add (by intro a b hab obtain ⟨a, rfl⟩ := hf a obtain ⟨b, rfl⟩ := hf b rw [← map_add] at hab exact - (isUnit_or_isUnit_of_isUnit_add <| IsLocalRingHom.map_nonunit _ hab).imp f.isUnit_map + (isUnit_or_isUnit_of_isUnit_add <| IsLocalHom.map_nonunit _ hab).imp f.isUnit_map f.isUnit_map) /-- If `f : R →+* S` is a surjective local ring hom, then the induced units map is surjective. -/ theorem surjective_units_map_of_local_ringHom [CommRing R] [CommRing S] (f : R →+* S) - (hf : Function.Surjective f) (h : IsLocalRingHom f) : + (hf : Function.Surjective f) (h : IsLocalHom f) : Function.Surjective (Units.map <| f.toMonoidHom) := by intro a obtain ⟨b, hb⟩ := hf (a : S) @@ -113,7 +128,7 @@ theorem surjective_units_map_of_local_ringHom [CommRing R] [CommRing S] (f : R /-- Every ring hom `f : K →+* R` from a division ring `K` to a nontrivial ring `R` is a local ring hom. -/ instance (priority := 100) {K R} [DivisionRing K] [CommRing R] [Nontrivial R] - (f : K →+* R) : IsLocalRingHom f where + (f : K →+* R) : IsLocalHom f where map_nonunit r hr := by simpa only [isUnit_iff_ne_zero, ne_eq, map_eq_zero] using hr.ne_zero end LocalRing diff --git a/Mathlib/RingTheory/Localization/AtPrime.lean b/Mathlib/RingTheory/Localization/AtPrime.lean index cec0198c34a92..90a63bac86990 100644 --- a/Mathlib/RingTheory/Localization/AtPrime.lean +++ b/Mathlib/RingTheory/Localization/AtPrime.lean @@ -218,14 +218,18 @@ theorem localRingHom_mk' (J : Ideal P) [J.IsPrime] (f : R →+* P) (hIJ : I = J. (⟨f y, le_comap_primeCompl_iff.mpr (ge_of_eq hIJ) y.2⟩ : J.primeCompl) := map_mk' _ _ _ -instance isLocalRingHom_localRingHom (J : Ideal P) [hJ : J.IsPrime] (f : R →+* P) - (hIJ : I = J.comap f) : IsLocalRingHom (localRingHom I J f hIJ) := - IsLocalRingHom.mk fun x hx => by +@[instance] +theorem isLocalHom_localRingHom (J : Ideal P) [hJ : J.IsPrime] (f : R →+* P) + (hIJ : I = J.comap f) : IsLocalHom (localRingHom I J f hIJ) := + IsLocalHom.mk fun x hx => by rcases IsLocalization.mk'_surjective I.primeCompl x with ⟨r, s, rfl⟩ rw [localRingHom_mk'] at hx rw [AtPrime.isUnit_mk'_iff] at hx ⊢ exact fun hr => hx ((SetLike.ext_iff.mp hIJ r).mp hr) +@[deprecated (since := "2024-10-10")] +alias isLocalRingHom_localRingHom := isLocalHom_localRingHom + theorem localRingHom_unique (J : Ideal P) [J.IsPrime] (f : R →+* P) (hIJ : I = J.comap f) {j : Localization.AtPrime I →+* Localization.AtPrime J} (hj : ∀ x : R, j (algebraMap _ _ x) = algebraMap _ _ (f x)) : localRingHom I J f hIJ = j := diff --git a/Mathlib/RingTheory/MvPowerSeries/Inverse.lean b/Mathlib/RingTheory/MvPowerSeries/Inverse.lean index d5160e8067589..ce619c30ce099 100644 --- a/Mathlib/RingTheory/MvPowerSeries/Inverse.lean +++ b/Mathlib/RingTheory/MvPowerSeries/Inverse.lean @@ -31,7 +31,7 @@ Instances are defined: * Formal power series over a local ring form a local ring. * The morphism `MvPowerSeries.map σ f : MvPowerSeries σ A →* MvPowerSeries σ B` - induced by a local morphism `f : A →+* B` (`IsLocalRingHom f`) + induced by a local morphism `f : A →+* B` (`IsLocalHom f`) of commutative rings is a *local* morphism. -/ @@ -167,13 +167,14 @@ end CommRing section LocalRing -variable {S : Type*} [CommRing R] [CommRing S] (f : R →+* S) [IsLocalRingHom f] +variable {S : Type*} [CommRing R] [CommRing S] (f : R →+* S) [IsLocalHom f] -- Thanks to the linter for informing us that this instance does -- not actually need R and S to be local rings! /-- The map between multivariate formal power series over the same indexing set induced by a local ring hom `A → B` is local -/ -instance map.isLocalRingHom : IsLocalRingHom (map σ f) := +@[instance] +theorem map.isLocalHom : IsLocalHom (map σ f) := ⟨by rintro φ ⟨ψ, h⟩ replace h := congr_arg (constantCoeff σ S) h @@ -183,6 +184,9 @@ instance map.isLocalRingHom : IsLocalRingHom (map σ f) := rcases isUnit_of_map_unit f _ this with ⟨c, hc⟩ exact isUnit_of_mul_eq_one φ (invOfUnit φ c) (mul_invOfUnit φ c hc.symm)⟩ +@[deprecated (since := "2024-10-10")] +alias map.isLocalRingHom := map.isLocalHom + end LocalRing section Field diff --git a/Mathlib/RingTheory/PowerSeries/Inverse.lean b/Mathlib/RingTheory/PowerSeries/Inverse.lean index aec805253fd3d..4a9202d3df4a7 100644 --- a/Mathlib/RingTheory/PowerSeries/Inverse.lean +++ b/Mathlib/RingTheory/PowerSeries/Inverse.lean @@ -266,10 +266,14 @@ end Field section LocalRing -variable {S : Type*} [CommRing R] [CommRing S] (f : R →+* S) [IsLocalRingHom f] +variable {S : Type*} [CommRing R] [CommRing S] (f : R →+* S) [IsLocalHom f] -instance map.isLocalRingHom : IsLocalRingHom (map f) := - MvPowerSeries.map.isLocalRingHom f +@[instance] +theorem map.isLocalHom : IsLocalHom (map f) := + MvPowerSeries.map.isLocalHom f + +@[deprecated (since := "2024-10-10")] +alias map.isLocalRingHom := map.isLocalHom variable [LocalRing R] [LocalRing S] diff --git a/Mathlib/RingTheory/SurjectiveOnStalks.lean b/Mathlib/RingTheory/SurjectiveOnStalks.lean index 94c5443047b71..c8c8f65ad3cba 100644 --- a/Mathlib/RingTheory/SurjectiveOnStalks.lean +++ b/Mathlib/RingTheory/SurjectiveOnStalks.lean @@ -174,7 +174,7 @@ lemma SurjectiveOnStalks.baseChange one_mul, mul_one, id_apply, ← e] rw [Algebra.algebraMap_eq_smul_one, ← smul_tmul', smul_mul_assoc] -lemma surjectiveOnStalks_iff_of_isLocalRingHom [LocalRing S] [IsLocalRingHom f] : +lemma surjectiveOnStalks_iff_of_isLocalHom [LocalRing S] [IsLocalHom f] : f.SurjectiveOnStalks ↔ Function.Surjective f := by refine ⟨fun H x ↦ ?_, fun h ↦ surjectiveOnStalks_of_surjective h⟩ obtain ⟨y, r, c, hc, hr, e⟩ := @@ -185,4 +185,7 @@ lemma surjectiveOnStalks_iff_of_isLocalRingHom [LocalRing S] [IsLocalRingHom f] apply hc.mul_right_injective simp only [← _root_.map_mul, ← mul_assoc, IsUnit.mul_val_inv, one_mul, e] +@[deprecated (since := "2024-10-10")] +alias surjectiveOnStalks_iff_of_isLocalRingHom := surjectiveOnStalks_iff_of_isLocalHom + end RingHom diff --git a/Mathlib/RingTheory/Valuation/ValExtension.lean b/Mathlib/RingTheory/Valuation/ValExtension.lean index 566aa0f741a77..11ee048743d9c 100644 --- a/Mathlib/RingTheory/Valuation/ValExtension.lean +++ b/Mathlib/RingTheory/Valuation/ValExtension.lean @@ -143,16 +143,20 @@ theorem algebraMap_injective [IsValExtension vK vA] [Nontrivial A] : ext apply RingHom.injective (algebraMap K A) h -instance instIsLocalRingHomValuationInteger {S ΓS: Type*} [CommRing S] +@[instance] +theorem instIsLocalHomValuationInteger {S ΓS: Type*} [CommRing S] [LinearOrderedCommGroupWithZero ΓS] - [Algebra R S] [IsLocalRingHom (algebraMap R S)] {vS : Valuation S ΓS} - [IsValExtension vR vS] : IsLocalRingHom (algebraMap vR.integer vS.integer) where + [Algebra R S] [IsLocalHom (algebraMap R S)] {vS : Valuation S ΓS} + [IsValExtension vR vS] : IsLocalHom (algebraMap vR.integer vS.integer) where map_nonunit r hr := by apply (Valuation.integer.integers (v := vR)).isUnit_of_one · exact (isUnit_map_iff (algebraMap R S) _).mp (hr.map (algebraMap _ S)) · apply (Valuation.integer.integers (v := vS)).one_of_isUnit at hr exact (val_map_eq_one_iff vR vS _).mp hr +@[deprecated (since := "2024-10-10")] +alias instIsLocalRingHomValuationInteger := instIsLocalHomValuationInteger + end integer end IsValExtension diff --git a/Mathlib/RingTheory/Valuation/ValuationSubring.lean b/Mathlib/RingTheory/Valuation/ValuationSubring.lean index 48d6df21160fa..d6d738435853c 100644 --- a/Mathlib/RingTheory/Valuation/ValuationSubring.lean +++ b/Mathlib/RingTheory/Valuation/ValuationSubring.lean @@ -646,7 +646,7 @@ theorem ker_unitGroupToResidueFieldUnits : theorem surjective_unitGroupToResidueFieldUnits : Function.Surjective A.unitGroupToResidueFieldUnits := (LocalRing.surjective_units_map_of_local_ringHom _ Ideal.Quotient.mk_surjective - LocalRing.isLocalRingHom_residue).comp + LocalRing.isLocalHom_residue).comp (MulEquiv.surjective _) /-- The quotient of the unit group of `A` by the principal unit group of `A` agrees with From 1f26636368232f82cf41fc3ab1793f6d61a9edc2 Mon Sep 17 00:00:00 2001 From: Moritz Firsching Date: Fri, 18 Oct 2024 08:51:11 +0000 Subject: [PATCH 128/505] chore(Algebra/Module/Equiv/Defs): remove `FunLike` instance (#17477) Co-authored-by: Moritz Firsching --- Mathlib/Algebra/Module/Equiv/Defs.lean | 9 --------- 1 file changed, 9 deletions(-) diff --git a/Mathlib/Algebra/Module/Equiv/Defs.lean b/Mathlib/Algebra/Module/Equiv/Defs.lean index 7d08836e2d79f..430a86911a3e0 100644 --- a/Mathlib/Algebra/Module/Equiv/Defs.lean +++ b/Mathlib/Algebra/Module/Equiv/Defs.lean @@ -168,15 +168,6 @@ instance : EquivLike (M ≃ₛₗ[σ] M₂) M M₂ where left_inv := LinearEquiv.left_inv right_inv := LinearEquiv.right_inv -/-- Helper instance for when inference gets stuck on following the normal chain -`EquivLike → FunLike`. - -TODO: this instance doesn't appear to be necessary: remove it (after benchmarking?) --/ -instance : FunLike (M ≃ₛₗ[σ] M₂) M M₂ where - coe := DFunLike.coe - coe_injective' := DFunLike.coe_injective - instance : SemilinearEquivClass (M ≃ₛₗ[σ] M₂) σ M M₂ where map_add := (·.map_add') --map_add' Porting note (#11215): TODO why did I need to change this? map_smulₛₗ := (·.map_smul') --map_smul' Porting note (#11215): TODO why did I need to change this? From 6dca5634f0e8d885d4dae292bfa910bb62a9e524 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 18 Oct 2024 08:51:13 +0000 Subject: [PATCH 129/505] chore(MeasureTheory): use `Disjoint` (#17890) --- .../ConditionalExpectation/CondexpL1.lean | 8 ++++---- Mathlib/MeasureTheory/Function/LpSpace.lean | 4 ++-- Mathlib/MeasureTheory/Integral/Bochner.lean | 9 ++++----- Mathlib/MeasureTheory/Integral/SetToL1.lean | 16 ++++++++-------- 4 files changed, 18 insertions(+), 19 deletions(-) diff --git a/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL1.lean b/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL1.lean index 633ff287c5581..45d4dfd32c71c 100644 --- a/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL1.lean +++ b/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL1.lean @@ -135,7 +135,7 @@ theorem norm_condexpIndL1Fin_le (hs : MeasurableSet s) (hμs : μ s ≠ ∞) (x exact lintegral_nnnorm_condexpIndSMul_le hm hs hμs x theorem condexpIndL1Fin_disjoint_union (hs : MeasurableSet s) (ht : MeasurableSet t) (hμs : μ s ≠ ∞) - (hμt : μ t ≠ ∞) (hst : s ∩ t = ∅) (x : G) : + (hμt : μ t ≠ ∞) (hst : Disjoint s t) (x : G) : condexpIndL1Fin hm (hs.union ht) ((measure_union_le s t).trans_lt (lt_top_iff_ne_top.mpr (ENNReal.add_ne_top.mpr ⟨hμs, hμt⟩))).ne x = condexpIndL1Fin hm hs hμs x + condexpIndL1Fin hm ht hμt x := by @@ -224,7 +224,7 @@ theorem continuous_condexpIndL1 : Continuous fun x : G => condexpIndL1 hm μ s x continuous_of_linear_of_bound condexpIndL1_add condexpIndL1_smul norm_condexpIndL1_le theorem condexpIndL1_disjoint_union (hs : MeasurableSet s) (ht : MeasurableSet t) (hμs : μ s ≠ ∞) - (hμt : μ t ≠ ∞) (hst : s ∩ t = ∅) (x : G) : + (hμt : μ t ≠ ∞) (hst : Disjoint s t) (x : G) : condexpIndL1 hm μ (s ∪ t) x = condexpIndL1 hm μ s x + condexpIndL1 hm μ t x := by have hμst : μ (s ∪ t) ≠ ∞ := ((measure_union_le s t).trans_lt (lt_top_iff_ne_top.mpr (ENNReal.add_ne_top.mpr ⟨hμs, hμt⟩))).ne @@ -283,12 +283,12 @@ theorem norm_condexpInd_le : ‖(condexpInd G hm μ s : G →L[ℝ] α →₁[μ ContinuousLinearMap.opNorm_le_bound _ ENNReal.toReal_nonneg norm_condexpInd_apply_le theorem condexpInd_disjoint_union_apply (hs : MeasurableSet s) (ht : MeasurableSet t) - (hμs : μ s ≠ ∞) (hμt : μ t ≠ ∞) (hst : s ∩ t = ∅) (x : G) : + (hμs : μ s ≠ ∞) (hμt : μ t ≠ ∞) (hst : Disjoint s t) (x : G) : condexpInd G hm μ (s ∪ t) x = condexpInd G hm μ s x + condexpInd G hm μ t x := condexpIndL1_disjoint_union hs ht hμs hμt hst x theorem condexpInd_disjoint_union (hs : MeasurableSet s) (ht : MeasurableSet t) (hμs : μ s ≠ ∞) - (hμt : μ t ≠ ∞) (hst : s ∩ t = ∅) : (condexpInd G hm μ (s ∪ t) : G →L[ℝ] α →₁[μ] G) = + (hμt : μ t ≠ ∞) (hst : Disjoint s t) : (condexpInd G hm μ (s ∪ t) : G →L[ℝ] α →₁[μ] G) = condexpInd G hm μ s + condexpInd G hm μ t := by ext1 x; push_cast; exact condexpInd_disjoint_union_apply hs ht hμs hμt hst x diff --git a/Mathlib/MeasureTheory/Function/LpSpace.lean b/Mathlib/MeasureTheory/Function/LpSpace.lean index a20a1c45824a7..7e5941d364518 100644 --- a/Mathlib/MeasureTheory/Function/LpSpace.lean +++ b/Mathlib/MeasureTheory/Function/LpSpace.lean @@ -857,7 +857,7 @@ theorem memℒp_add_of_disjoint {f g : α → E} (h : Disjoint (support f) (supp /-- The indicator of a disjoint union of two sets is the sum of the indicators of the sets. -/ theorem indicatorConstLp_disjoint_union {s t : Set α} (hs : MeasurableSet s) (ht : MeasurableSet t) - (hμs : μ s ≠ ∞) (hμt : μ t ≠ ∞) (hst : s ∩ t = ∅) (c : E) : + (hμs : μ s ≠ ∞) (hμt : μ t ≠ ∞) (hst : Disjoint s t) (c : E) : indicatorConstLp p (hs.union ht) (measure_union_ne_top hμs hμt) c = indicatorConstLp p hs hμs c + indicatorConstLp p ht hμt c := by ext1 @@ -865,7 +865,7 @@ theorem indicatorConstLp_disjoint_union {s t : Set α} (hs : MeasurableSet s) (h refine EventuallyEq.trans ?_ (EventuallyEq.add indicatorConstLp_coeFn.symm indicatorConstLp_coeFn.symm) - rw [Set.indicator_union_of_disjoint (Set.disjoint_iff_inter_eq_empty.mpr hst) _] + rw [Set.indicator_union_of_disjoint hst] end IndicatorConstLp diff --git a/Mathlib/MeasureTheory/Integral/Bochner.lean b/Mathlib/MeasureTheory/Integral/Bochner.lean index b4eab8d959db6..b497a13fc272b 100644 --- a/Mathlib/MeasureTheory/Integral/Bochner.lean +++ b/Mathlib/MeasureTheory/Integral/Bochner.lean @@ -201,18 +201,17 @@ theorem weightedSMul_null {s : Set α} (h_zero : μ s = 0) : (weightedSMul μ s ext1 x; rw [weightedSMul_apply, h_zero]; simp theorem weightedSMul_union' (s t : Set α) (ht : MeasurableSet t) (hs_finite : μ s ≠ ∞) - (ht_finite : μ t ≠ ∞) (h_inter : s ∩ t = ∅) : + (ht_finite : μ t ≠ ∞) (hdisj : Disjoint s t) : (weightedSMul μ (s ∪ t) : F →L[ℝ] F) = weightedSMul μ s + weightedSMul μ t := by ext1 x - simp_rw [add_apply, weightedSMul_apply, - measure_union (Set.disjoint_iff_inter_eq_empty.mpr h_inter) ht, + simp_rw [add_apply, weightedSMul_apply, measure_union hdisj ht, ENNReal.toReal_add hs_finite ht_finite, add_smul] @[nolint unusedArguments] theorem weightedSMul_union (s t : Set α) (_hs : MeasurableSet s) (ht : MeasurableSet t) - (hs_finite : μ s ≠ ∞) (ht_finite : μ t ≠ ∞) (h_inter : s ∩ t = ∅) : + (hs_finite : μ s ≠ ∞) (ht_finite : μ t ≠ ∞) (hdisj : Disjoint s t) : (weightedSMul μ (s ∪ t) : F →L[ℝ] F) = weightedSMul μ s + weightedSMul μ t := - weightedSMul_union' s t ht hs_finite ht_finite h_inter + weightedSMul_union' s t ht hs_finite ht_finite hdisj theorem weightedSMul_smul [NormedField 𝕜] [NormedSpace 𝕜 F] [SMulCommClass ℝ 𝕜 F] (c : 𝕜) (s : Set α) (x : F) : weightedSMul μ s (c • x) = c • weightedSMul μ s x := by diff --git a/Mathlib/MeasureTheory/Integral/SetToL1.lean b/Mathlib/MeasureTheory/Integral/SetToL1.lean index e4fe8a61b787a..3d16c95fed2c1 100644 --- a/Mathlib/MeasureTheory/Integral/SetToL1.lean +++ b/Mathlib/MeasureTheory/Integral/SetToL1.lean @@ -9,7 +9,7 @@ import Mathlib.MeasureTheory.Function.SimpleFuncDenseLp # Extension of a linear function from indicators to L1 Let `T : Set α → E →L[ℝ] F` be additive for measurable sets with finite measure, in the sense that -for `s, t` two such sets, `s ∩ t = ∅ → T (s ∪ t) = T s + T t`. `T` is akin to a bilinear map on +for `s, t` two such sets, `Disjoint s t → T (s ∪ t) = T s + T t`. `T` is akin to a bilinear map on `Set α × E`, or a linear map on indicator functions. This file constructs an extension of `T` to integrable simple functions, which are finite sums of @@ -23,7 +23,7 @@ expectation of an integrable function in `MeasureTheory.Function.ConditionalExpe ## Main Definitions - `FinMeasAdditive μ T`: the property that `T` is additive on measurable sets with finite measure. - For two such sets, `s ∩ t = ∅ → T (s ∪ t) = T s + T t`. + For two such sets, `Disjoint s t → T (s ∪ t) = T s + T t`. - `DominatedFinMeasAdditive μ T C`: `FinMeasAdditive μ T ∧ ∀ s, ‖T s‖ ≤ C * (μ s).toReal`. This is the property needed to perform the extension from indicators to L1. - `setToL1 (hT : DominatedFinMeasAdditive μ T C) : (α →₁[μ] E) →L[ℝ] F`: the extension of `T` @@ -90,7 +90,8 @@ section FinMeasAdditive sets with finite measure is the sum of its values on each set. -/ def FinMeasAdditive {β} [AddMonoid β] {_ : MeasurableSpace α} (μ : Measure α) (T : Set α → β) : Prop := - ∀ s t, MeasurableSet s → MeasurableSet t → μ s ≠ ∞ → μ t ≠ ∞ → s ∩ t = ∅ → T (s ∪ t) = T s + T t + ∀ s t, MeasurableSet s → MeasurableSet t → μ s ≠ ∞ → μ t ≠ ∞ → Disjoint s t → + T (s ∪ t) = T s + T t namespace FinMeasAdditive @@ -133,7 +134,7 @@ theorem smul_measure_iff (c : ℝ≥0∞) (hc_ne_zero : c ≠ 0) (hc_ne_top : c theorem map_empty_eq_zero {β} [AddCancelMonoid β] {T : Set α → β} (hT : FinMeasAdditive μ T) : T ∅ = 0 := by have h_empty : μ ∅ ≠ ∞ := (measure_empty.le.trans_lt ENNReal.coe_lt_top).ne - specialize hT ∅ ∅ MeasurableSet.empty MeasurableSet.empty h_empty h_empty (Set.inter_empty ∅) + specialize hT ∅ ∅ MeasurableSet.empty MeasurableSet.empty h_empty h_empty (disjoint_empty _) rw [Set.union_empty] at hT nth_rw 1 [← add_zero (T ∅)] at hT exact (add_left_cancel hT).symm @@ -160,10 +161,9 @@ theorem map_iUnion_fin_meas_set_eq_sum (T : Set α → β) (T_empty : T ∅ = 0) · congr; convert Finset.iSup_insert a s S · exact (measure_biUnion_lt_top s.finite_toSet fun i hi ↦ (hps i <| Finset.mem_insert_of_mem hi).lt_top).ne - · simp_rw [Set.inter_iUnion] - refine iUnion_eq_empty.mpr fun i => iUnion_eq_empty.mpr fun hi => ?_ - rw [← Set.disjoint_iff_inter_eq_empty] - refine h_disj a (Finset.mem_insert_self a s) i (Finset.mem_insert_of_mem hi) fun hai => ?_ + · simp_rw [Set.disjoint_iUnion_right] + intro i hi + refine h_disj a (Finset.mem_insert_self a s) i (Finset.mem_insert_of_mem hi) fun hai ↦ ?_ rw [← hai] at hi exact has hi From ec349bacca2d1e67eae5142d16b30a202510578d Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 18 Oct 2024 09:18:36 +0000 Subject: [PATCH 130/505] chore(PosPart): Make more lemmas simp (#17053) --- Mathlib/Algebra/Order/Group/PosPart.lean | 9 +++++++-- 1 file changed, 7 insertions(+), 2 deletions(-) diff --git a/Mathlib/Algebra/Order/Group/PosPart.lean b/Mathlib/Algebra/Order/Group/PosPart.lean index 085e56c1b0be2..eb8835fa80134 100644 --- a/Mathlib/Algebra/Order/Group/PosPart.lean +++ b/Mathlib/Algebra/Order/Group/PosPart.lean @@ -65,7 +65,7 @@ instance instLeOnePart : LeOnePart α where @[to_additive] lemma oneLePart_mono : Monotone (·⁺ᵐ : α → α) := fun _a _b hab ↦ sup_le_sup_right hab _ -@[to_additive (attr := simp)] lemma oneLePart_one : (1 : α)⁺ᵐ = 1 := sup_idem _ +@[to_additive (attr := simp high)] lemma oneLePart_one : (1 : α)⁺ᵐ = 1 := sup_idem _ @[to_additive (attr := simp)] lemma leOnePart_one : (1 : α)⁻ᵐ = 1 := by simp [leOnePart] @@ -79,8 +79,10 @@ instance instLeOnePart : LeOnePart α where @[to_additive] lemma inv_le_leOnePart (a : α) : a⁻¹ ≤ a⁻ᵐ := le_sup_left @[to_additive (attr := simp)] lemma oneLePart_eq_self : a⁺ᵐ = a ↔ 1 ≤ a := sup_eq_left +@[to_additive (attr := simp)] lemma oneLePart_eq_one : a⁺ᵐ = 1 ↔ a ≤ 1 := sup_eq_right -@[to_additive] lemma oneLePart_eq_one : a⁺ᵐ = 1 ↔ a ≤ 1 := sup_eq_right +@[to_additive (attr := simp)] alias ⟨_, oneLePart_of_one_le⟩ := oneLePart_eq_self +@[to_additive (attr := simp)] alias ⟨_, oneLePart_of_le_one⟩ := oneLePart_eq_one /-- See also `leOnePart_eq_inv`. -/ @[to_additive "See also `negPart_eq_neg`."] @@ -114,6 +116,9 @@ variable [CovariantClass α α (· * ·) (· ≤ ·)] @[to_additive (attr := simp)] lemma leOnePart_eq_one : a⁻ᵐ = 1 ↔ 1 ≤ a := by simp [leOnePart_eq_one'] +@[to_additive (attr := simp)] alias ⟨_, leOnePart_of_le_one⟩ := leOnePart_eq_inv +@[to_additive (attr := simp)] alias ⟨_, leOnePart_of_one_le⟩ := leOnePart_eq_one + @[to_additive (attr := simp) negPart_pos] lemma one_lt_ltOnePart (ha : a < 1) : 1 < a⁻ᵐ := by rwa [leOnePart_eq_inv.2 ha.le, one_lt_inv'] From 3b5cedb7b2da9faae67402d7a66a75e5f44b666e Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 18 Oct 2024 09:18:37 +0000 Subject: [PATCH 131/505] feat(Topology): define `ContinuousEval{,Const}` classes (#17319) --- Mathlib.lean | 2 + .../NormedSpace/Multilinear/Basic.lean | 27 +++--- Mathlib/Analysis/ODE/PicardLindelof.lean | 4 +- .../Topology/Algebra/ContinuousMonoidHom.lean | 15 +++- .../Algebra/Module/Alternating/Topology.lean | 26 +++--- .../Algebra/Module/Multilinear/Topology.lean | 27 +++--- .../Algebra/Module/StrongTopology.lean | 30 ++++--- Mathlib/Topology/CompactOpen.lean | 41 +++++---- Mathlib/Topology/Connected/PathConnected.lean | 14 +-- Mathlib/Topology/ContinuousMap/Algebra.lean | 40 +++++---- .../ContinuousMap/ContinuousMapZero.lean | 12 ++- Mathlib/Topology/Hom/ContinuousEval.lean | 69 ++++++++++++++ Mathlib/Topology/Hom/ContinuousEvalConst.lean | 89 +++++++++++++++++++ Mathlib/Topology/Homotopy/HSpaces.lean | 2 +- Mathlib/Topology/Homotopy/HomotopyGroup.lean | 8 +- Mathlib/Topology/Homotopy/Path.lean | 10 +-- 16 files changed, 307 insertions(+), 109 deletions(-) create mode 100644 Mathlib/Topology/Hom/ContinuousEval.lean create mode 100644 Mathlib/Topology/Hom/ContinuousEvalConst.lean diff --git a/Mathlib.lean b/Mathlib.lean index 6b0d74489e5e7..1e5548b290acc 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4716,6 +4716,8 @@ import Mathlib.Topology.Filter import Mathlib.Topology.GDelta import Mathlib.Topology.Germ import Mathlib.Topology.Gluing +import Mathlib.Topology.Hom.ContinuousEval +import Mathlib.Topology.Hom.ContinuousEvalConst import Mathlib.Topology.Hom.Open import Mathlib.Topology.Homeomorph import Mathlib.Topology.Homotopy.Basic diff --git a/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean b/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean index ceb978b4e25e7..7f88338734d62 100644 --- a/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean +++ b/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean @@ -80,16 +80,19 @@ variable {𝕜 ι : Type*} {E : ι → Type*} {F : Type*} [NormedField 𝕜] [Finite ι] [∀ i, SeminormedAddCommGroup (E i)] [∀ i, NormedSpace 𝕜 (E i)] [TopologicalSpace F] [AddCommGroup F] [TopologicalAddGroup F] [Module 𝕜 F] -/-- Applying a multilinear map to a vector is continuous in both coordinates. -/ -theorem ContinuousMultilinearMap.continuous_eval : - Continuous fun p : ContinuousMultilinearMap 𝕜 E F × ∀ i, E i => p.1 p.2 := by - cases nonempty_fintype ι - let _ := TopologicalAddGroup.toUniformSpace F - have := comm_topologicalAddGroup_is_uniform (G := F) - refine (UniformOnFun.continuousOn_eval₂ fun m ↦ ?_).comp_continuous - (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⟩ +instance ContinuousMultilinearMap.instContinuousEval : + ContinuousEval (ContinuousMultilinearMap 𝕜 E F) (Π i, E i) F where + continuous_eval := by + cases nonempty_fintype ι + let _ := TopologicalAddGroup.toUniformSpace F + have := comm_topologicalAddGroup_is_uniform (G := F) + refine (UniformOnFun.continuousOn_eval₂ fun m ↦ ?_).comp_continuous + (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⟩ + +@[deprecated (since := "2024-10-05")] +protected alias ContinuousMultilinearMap.continuous_eval := continuous_eval namespace ContinuousLinearMap @@ -97,8 +100,8 @@ variable {G : Type*} [AddCommGroup G] [TopologicalSpace G] [Module 𝕜 G] [Cont (f : G →L[𝕜] ContinuousMultilinearMap 𝕜 E F) lemma continuous_uncurry_of_multilinear : - Continuous (fun (p : G × (Π i, E i)) ↦ f p.1 p.2) := - ContinuousMultilinearMap.continuous_eval.comp <| .prodMap (map_continuous f) continuous_id + Continuous (fun (p : G × (Π i, E i)) ↦ f p.1 p.2) := by + fun_prop lemma continuousOn_uncurry_of_multilinear {s} : ContinuousOn (fun (p : G × (Π i, E i)) ↦ f p.1 p.2) s := diff --git a/Mathlib/Analysis/ODE/PicardLindelof.lean b/Mathlib/Analysis/ODE/PicardLindelof.lean index 1f6996fcbf20e..bf6d91625d9e7 100644 --- a/Mathlib/Analysis/ODE/PicardLindelof.lean +++ b/Mathlib/Analysis/ODE/PicardLindelof.lean @@ -223,10 +223,10 @@ instance [CompleteSpace E] : CompleteSpace v.FunSpace := by refine (completeSpace_iff_isComplete_range isUniformInducing_toContinuousMap).2 (IsClosed.isComplete ?_) rw [range_toContinuousMap, setOf_and] - refine (isClosed_eq (ContinuousMap.continuous_eval_const _) continuous_const).inter ?_ + refine (isClosed_eq (continuous_eval_const _) continuous_const).inter ?_ have : IsClosed {f : Icc v.tMin v.tMax → E | LipschitzWith v.C f} := isClosed_setOf_lipschitzWith v.C - exact this.preimage ContinuousMap.continuous_coe + exact this.preimage continuous_coeFun theorem intervalIntegrable_vComp (t₁ t₂ : ℝ) : IntervalIntegrable f.vComp volume t₁ t₂ := f.continuous_vComp.intervalIntegrable _ _ diff --git a/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean b/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean index 4e852f064519b..a41b781ceaadc 100644 --- a/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean +++ b/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean @@ -240,6 +240,15 @@ theorem embedding_toContinuousMap : Embedding (toContinuousMap : ContinuousMonoidHom A B → C(A, B)) := ⟨inducing_toContinuousMap A B, toContinuousMap_injective⟩ +@[to_additive] +instance instContinuousEvalConst : ContinuousEvalConst (ContinuousMonoidHom A B) A B := + .of_continuous_forget (inducing_toContinuousMap A B).continuous + +@[to_additive] +instance instContinuousEval [LocallyCompactPair A B] : + ContinuousEval (ContinuousMonoidHom A B) A B := + .of_continuous_forget (inducing_toContinuousMap A B).continuous + @[to_additive] lemma range_toContinuousMap : Set.range (toContinuousMap : ContinuousMonoidHom A B → C(A, B)) = @@ -254,10 +263,10 @@ theorem closedEmbedding_toContinuousMap [ContinuousMul B] [T2Space B] : toEmbedding := embedding_toContinuousMap A B isClosed_range := by simp only [range_toContinuousMap, Set.setOf_and, Set.setOf_forall] - refine .inter (isClosed_singleton.preimage (ContinuousMap.continuous_eval_const 1)) <| + refine .inter (isClosed_singleton.preimage (continuous_eval_const 1)) <| isClosed_iInter fun x ↦ isClosed_iInter fun y ↦ ?_ - exact isClosed_eq (ContinuousMap.continuous_eval_const (x * y)) <| - .mul (ContinuousMap.continuous_eval_const x) (ContinuousMap.continuous_eval_const y) + exact isClosed_eq (continuous_eval_const (x * y)) <| + .mul (continuous_eval_const x) (continuous_eval_const y) variable {A B C D E} diff --git a/Mathlib/Topology/Algebra/Module/Alternating/Topology.lean b/Mathlib/Topology/Algebra/Module/Alternating/Topology.lean index 32efbbd5a39c7..52bcb32bef151 100644 --- a/Mathlib/Topology/Algebra/Module/Alternating/Topology.lean +++ b/Mathlib/Topology/Algebra/Module/Alternating/Topology.lean @@ -36,7 +36,7 @@ lemma isClosed_range_toContinuousMultilinearMap [ContinuousSMul 𝕜 E] [T2Space ContinuousMultilinearMap 𝕜 (fun _ : ι ↦ E) F)) := by simp only [range_toContinuousMultilinearMap, setOf_forall] repeat refine isClosed_iInter fun _ ↦ ?_ - exact isClosed_singleton.preimage (ContinuousMultilinearMap.continuous_eval_const _) + exact isClosed_singleton.preimage (continuous_eval_const _) end IsClosedRange @@ -125,6 +125,10 @@ lemma embedding_toContinuousMultilinearMap : haveI := comm_topologicalAddGroup_is_uniform (G := F) isUniformEmbedding_toContinuousMultilinearMap.embedding +instance instTopologicalAddGroup : TopologicalAddGroup (E [⋀^ι]→L[𝕜] F) := + embedding_toContinuousMultilinearMap.topologicalAddGroup + (toContinuousMultilinearMapLinear (R := ℕ)) + @[continuity, fun_prop] lemma continuous_toContinuousMultilinearMap : Continuous (toContinuousMultilinearMap : (E [⋀^ι]→L[𝕜] F → _)) := @@ -159,21 +163,19 @@ lemma closedEmbedding_toContinuousMultilinearMap [T2Space F] : (E [⋀^ι]→L[𝕜] F) → ContinuousMultilinearMap 𝕜 (fun _ : ι ↦ E) F) := ⟨embedding_toContinuousMultilinearMap, isClosed_range_toContinuousMultilinearMap⟩ -@[continuity, fun_prop] -theorem continuous_eval_const (x : ι → E) : - Continuous fun p : E [⋀^ι]→L[𝕜] F ↦ p x := - (ContinuousMultilinearMap.continuous_eval_const x).comp continuous_toContinuousMultilinearMap +instance instContinuousEvalConst : ContinuousEvalConst (E [⋀^ι]→L[𝕜] F) (ι → E) F := + .of_continuous_forget continuous_toContinuousMultilinearMap -theorem continuous_coe_fun : - Continuous (DFunLike.coe : E [⋀^ι]→L[𝕜] F → (ι → E) → F) := - continuous_pi continuous_eval_const +@[deprecated (since := "2024-10-05")] +protected alias continuous_eval_const := continuous_eval_const + +@[deprecated (since := "2024-10-05")] +protected alias continuous_coe_fun := continuous_coeFun instance instT2Space [T2Space F] : T2Space (E [⋀^ι]→L[𝕜] F) := - .of_injective_continuous DFunLike.coe_injective continuous_coe_fun + .of_injective_continuous DFunLike.coe_injective continuous_coeFun -instance instT3Space [T2Space F] : T2Space (E [⋀^ι]→L[𝕜] F) := - letI : UniformSpace F := TopologicalAddGroup.toUniformSpace F - haveI : UniformAddGroup F := comm_topologicalAddGroup_is_uniform +instance instT3Space [T2Space F] : T3Space (E [⋀^ι]→L[𝕜] F) := inferInstance section RestrictScalars diff --git a/Mathlib/Topology/Algebra/Module/Multilinear/Topology.lean b/Mathlib/Topology/Algebra/Module/Multilinear/Topology.lean index baf4b78201f9a..19af25db016c5 100644 --- a/Mathlib/Topology/Algebra/Module/Multilinear/Topology.lean +++ b/Mathlib/Topology/Algebra/Module/Multilinear/Topology.lean @@ -5,6 +5,7 @@ Authors: Yury Kudryashov -/ import Mathlib.Topology.Algebra.Module.Multilinear.Bounded import Mathlib.Topology.Algebra.Module.UniformConvergence +import Mathlib.Topology.Hom.ContinuousEvalConst /-! # Topology on continuous multilinear maps @@ -152,6 +153,11 @@ end UniformAddGroup variable [TopologicalSpace F] [TopologicalAddGroup F] +instance instTopologicalAddGroup : TopologicalAddGroup (ContinuousMultilinearMap 𝕜 E F) := + letI := TopologicalAddGroup.toUniformSpace F + haveI := comm_topologicalAddGroup_is_uniform (G := F) + inferInstance + instance instContinuousConstSMul {M : Type*} [Monoid M] [DistribMulAction M F] [SMulCommClass 𝕜 M F] [ContinuousConstSMul M F] : ContinuousConstSMul M (ContinuousMultilinearMap 𝕜 E F) := by @@ -188,20 +194,21 @@ theorem hasBasis_nhds_zero : variable [∀ i, ContinuousSMul 𝕜 (E i)] -theorem continuous_eval_const (x : Π i, E i) : - Continuous fun p : ContinuousMultilinearMap 𝕜 E F ↦ p x := by - letI := TopologicalAddGroup.toUniformSpace F - haveI := comm_topologicalAddGroup_is_uniform (G := F) - exact (uniformContinuous_eval_const x).continuous +instance : ContinuousEvalConst (ContinuousMultilinearMap 𝕜 E F) (Π i, E i) F where + continuous_eval_const x := + let _ := TopologicalAddGroup.toUniformSpace F + have _ := comm_topologicalAddGroup_is_uniform (G := F) + (uniformContinuous_eval_const x).continuous +@[deprecated (since := "2024-10-05")] protected alias continuous_eval_const := continuous_eval_const @[deprecated (since := "2024-04-10")] alias continuous_eval_left := continuous_eval_const - -theorem continuous_coe_fun : - Continuous (DFunLike.coe : ContinuousMultilinearMap 𝕜 E F → (Π i, E i) → F) := - continuous_pi continuous_eval_const +@[deprecated (since := "2024-10-05")] protected alias continuous_coe_fun := continuous_coeFun instance instT2Space [T2Space F] : T2Space (ContinuousMultilinearMap 𝕜 E F) := - .of_injective_continuous DFunLike.coe_injective continuous_coe_fun + .of_injective_continuous DFunLike.coe_injective continuous_coeFun + +instance instT3Space [T2Space F] : T3Space (ContinuousMultilinearMap 𝕜 E F) := + inferInstance section RestrictScalars diff --git a/Mathlib/Topology/Algebra/Module/StrongTopology.lean b/Mathlib/Topology/Algebra/Module/StrongTopology.lean index 8e988265fd745..3132c5f3d16b8 100644 --- a/Mathlib/Topology/Algebra/Module/StrongTopology.lean +++ b/Mathlib/Topology/Algebra/Module/StrongTopology.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Anatole Dedecker -/ import Mathlib.Topology.Algebra.Module.UniformConvergence +import Mathlib.Topology.Hom.ContinuousEvalConst /-! # Strong topologies on the space of continuous linear maps @@ -151,12 +152,19 @@ instance instTopologicalAddGroup [TopologicalSpace F] [TopologicalAddGroup F] haveI : UniformAddGroup F := comm_topologicalAddGroup_is_uniform infer_instance +theorem continuousEvalConst [TopologicalSpace F] [TopologicalAddGroup F] + (𝔖 : Set (Set E)) (h𝔖 : ⋃₀ 𝔖 = Set.univ) : + ContinuousEvalConst (UniformConvergenceCLM σ F 𝔖) E F where + continuous_eval_const x := by + letI : UniformSpace F := TopologicalAddGroup.toUniformSpace F + haveI : UniformAddGroup F := comm_topologicalAddGroup_is_uniform + exact (UniformOnFun.uniformContinuous_eval h𝔖 x).continuous.comp + (embedding_coeFn σ F 𝔖).continuous + theorem t2Space [TopologicalSpace F] [TopologicalAddGroup F] [T2Space F] - (𝔖 : Set (Set E)) (h𝔖 : ⋃₀ 𝔖 = univ) : T2Space (UniformConvergenceCLM σ F 𝔖) := by - letI : UniformSpace F := TopologicalAddGroup.toUniformSpace F - haveI : UniformAddGroup F := comm_topologicalAddGroup_is_uniform - haveI : T2Space (E →ᵤ[𝔖] F) := UniformOnFun.t2Space_of_covering h𝔖 - exact (embedding_coeFn σ F 𝔖).t2Space + (𝔖 : Set (Set E)) (h𝔖 : ⋃₀ 𝔖 = univ) : T2Space (UniformConvergenceCLM σ F 𝔖) := + have := continuousEvalConst σ F 𝔖 h𝔖 + .of_injective_continuous DFunLike.coe_injective continuous_coeFun instance instDistribMulAction (M : Type*) [Monoid M] [DistribMulAction M F] [SMulCommClass 𝕜₂ M F] [TopologicalSpace F] [TopologicalAddGroup F] [ContinuousConstSMul M F] (𝔖 : Set (Set E)) : @@ -321,11 +329,13 @@ instance uniformSpace [UniformSpace F] [UniformAddGroup F] : UniformSpace (E → instance uniformAddGroup [UniformSpace F] [UniformAddGroup F] : UniformAddGroup (E →SL[σ] F) := UniformConvergenceCLM.instUniformAddGroup σ F _ -instance [TopologicalSpace F] [TopologicalAddGroup F] [ContinuousSMul 𝕜₁ E] [T2Space F] : - T2Space (E →SL[σ] F) := - UniformConvergenceCLM.t2Space σ F _ - (Set.eq_univ_of_forall fun x => - Set.mem_sUnion_of_mem (Set.mem_singleton x) (isVonNBounded_singleton x)) +instance instContinuousEvalConst [TopologicalSpace F] [TopologicalAddGroup F] + [ContinuousSMul 𝕜₁ E] : ContinuousEvalConst (E →SL[σ] F) E F := + UniformConvergenceCLM.continuousEvalConst σ F _ Bornology.isVonNBounded_covers + +instance instT2Space [TopologicalSpace F] [TopologicalAddGroup F] [ContinuousSMul 𝕜₁ E] + [T2Space F] : T2Space (E →SL[σ] F) := + UniformConvergenceCLM.t2Space σ F _ Bornology.isVonNBounded_covers protected theorem hasBasis_nhds_zero_of_basis [TopologicalSpace F] [TopologicalAddGroup F] {ι : Type*} {p : ι → Prop} {b : ι → Set F} (h : (𝓝 0 : Filter F).HasBasis p b) : diff --git a/Mathlib/Topology/CompactOpen.lean b/Mathlib/Topology/CompactOpen.lean index c4e0af021885c..9c199df740de2 100644 --- a/Mathlib/Topology/CompactOpen.lean +++ b/Mathlib/Topology/CompactOpen.lean @@ -3,6 +3,7 @@ Copyright (c) 2018 Reid Barton. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Reid Barton -/ +import Mathlib.Topology.Hom.ContinuousEval import Mathlib.Topology.ContinuousMap.Basic /-! @@ -160,27 +161,24 @@ section Ev /-- The evaluation map `C(X, Y) × X → Y` is continuous if `X, Y` is a locally compact pair of spaces. -/ -@[continuity] -theorem continuous_eval [LocallyCompactPair X Y] : Continuous fun p : C(X, Y) × X => p.1 p.2 := by - simp_rw [continuous_iff_continuousAt, ContinuousAt, (nhds_basis_opens _).tendsto_right_iff] - rintro ⟨f, x⟩ U ⟨hx : f x ∈ U, hU : IsOpen U⟩ - rcases exists_mem_nhds_isCompact_mapsTo f.continuous (hU.mem_nhds hx) with ⟨K, hxK, hK, hKU⟩ - filter_upwards [prod_mem_nhds (eventually_mapsTo hK hU hKU) hxK] using fun _ h ↦ h.1 h.2 +instance [LocallyCompactPair X Y] : ContinuousEval C(X, Y) X Y where + continuous_eval := by + simp_rw [continuous_iff_continuousAt, ContinuousAt, (nhds_basis_opens _).tendsto_right_iff] + rintro ⟨f, x⟩ U ⟨hx : f x ∈ U, hU : IsOpen U⟩ + rcases exists_mem_nhds_isCompact_mapsTo f.continuous (hU.mem_nhds hx) with ⟨K, hxK, hK, hKU⟩ + filter_upwards [prod_mem_nhds (eventually_mapsTo hK hU hKU) hxK] using fun _ h ↦ h.1 h.2 -/-- Evaluation of a continuous map `f` at a point `x` is continuous in `f`. +@[deprecated (since := "2024-10-01")] protected alias continuous_eval := continuous_eval -Porting note: merged `continuous_eval_const` with `continuous_eval_const'` removing unneeded -assumptions. -/ -@[continuity] -theorem continuous_eval_const (a : X) : Continuous fun f : C(X, Y) => f a := - continuous_def.2 fun U hU ↦ by simpa using isOpen_setOf_mapsTo (isCompact_singleton (x := a)) hU +instance : ContinuousEvalConst C(X, Y) X Y where + continuous_eval_const x := + continuous_def.2 fun U hU ↦ by simpa using isOpen_setOf_mapsTo isCompact_singleton hU -/-- Coercion from `C(X, Y)` with compact-open topology to `X → Y` with pointwise convergence -topology is a continuous map. +@[deprecated (since := "2024-10-01")] protected alias continuous_eval_const := continuous_eval_const -Porting note: merged `continuous_coe` with `continuous_coe'` removing unneeded assumptions. -/ +@[deprecated continuous_coeFun (since := "2024-10-01")] theorem continuous_coe : Continuous ((⇑) : C(X, Y) → (X → Y)) := - continuous_pi continuous_eval_const + continuous_coeFun lemma isClosed_setOf_mapsTo {t : Set Y} (ht : IsClosed t) (s : Set X) : IsClosed {f : C(X, Y) | MapsTo f s t} := @@ -192,7 +190,7 @@ lemma isClopen_setOf_mapsTo (hK : IsCompact K) (hU : IsClopen U) : @[norm_cast] lemma specializes_coe {f g : C(X, Y)} : ⇑f ⤳ ⇑g ↔ f ⤳ g := by - refine ⟨fun h ↦ ?_, fun h ↦ h.map continuous_coe⟩ + refine ⟨fun h ↦ ?_, fun h ↦ h.map continuous_coeFun⟩ suffices ∀ K, IsCompact K → ∀ U, IsOpen U → MapsTo g K U → MapsTo f K U by simpa [specializes_iff_pure, nhds_compactOpen] exact fun K _ U hU hg x hx ↦ (h.map (continuous_apply x)).mem_open hU (hg hx) @@ -202,7 +200,7 @@ lemma inseparable_coe {f g : C(X, Y)} : Inseparable (f : X → Y) g ↔ Insepara simp only [inseparable_iff_specializes_and, specializes_coe] instance [T0Space Y] : T0Space C(X, Y) := - t0Space_of_injective_of_continuous DFunLike.coe_injective continuous_coe + t0Space_of_injective_of_continuous DFunLike.coe_injective continuous_coeFun instance [R0Space Y] : R0Space C(X, Y) where specializes_symmetric f g h := by @@ -210,10 +208,10 @@ instance [R0Space Y] : R0Space C(X, Y) where exact h.symm instance [T1Space Y] : T1Space C(X, Y) := - t1Space_of_injective_of_continuous DFunLike.coe_injective continuous_coe + t1Space_of_injective_of_continuous DFunLike.coe_injective continuous_coeFun instance [R1Space Y] : R1Space C(X, Y) := - .of_continuous_specializes_imp continuous_coe fun _ _ ↦ specializes_coe.1 + .of_continuous_specializes_imp continuous_coeFun fun _ _ ↦ specializes_coe.1 instance [T2Space Y] : T2Space C(X, Y) := inferInstance @@ -388,7 +386,8 @@ 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.prodMap continuous_id) + dsimp [Function.comp_def] + exact (continuous_fst.fst.eval continuous_fst.snd).eval continuous_snd /-- 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 e5798309d30d4..4cc61d6513e73 100644 --- a/Mathlib/Topology/Connected/PathConnected.lean +++ b/Mathlib/Topology/Connected/PathConnected.lean @@ -198,14 +198,14 @@ compact-open topology on the space `C(I,X)` of continuous maps from `I` to `X`. instance topologicalSpace : TopologicalSpace (Path x y) := TopologicalSpace.induced ((↑) : _ → C(I, X)) ContinuousMap.compactOpen -theorem continuous_eval : Continuous fun p : Path x y × I => p.1 p.2 := - ContinuousMap.continuous_eval.comp <| - (continuous_induced_dom (α := Path x y)).prodMap continuous_id +instance : ContinuousEval (Path x y) I X := .of_continuous_forget continuous_induced_dom -@[continuity] +@[deprecated (since := "2024-10-04")] protected alias continuous_eval := continuous_eval + +@[deprecated Continuous.eval (since := "2024-10-04")] theorem _root_.Continuous.path_eval {Y} [TopologicalSpace Y] {f : Y → Path x y} {g : Y → I} - (hf : Continuous f) (hg : Continuous g) : Continuous fun y => f y (g y) := - Continuous.comp continuous_eval (hf.prod_mk hg) + (hf : Continuous f) (hg : Continuous g) : Continuous fun y => f y (g y) := by + continuity theorem continuous_uncurry_iff {Y} [TopologicalSpace Y] {g : Y → Path x y} : Continuous ↿g ↔ Continuous g := @@ -427,7 +427,7 @@ theorem symm_continuous_family {ι : Type*} [TopologicalSpace ι] @[continuity] theorem continuous_symm : Continuous (symm : Path x y → Path y x) := - continuous_uncurry_iff.mp <| symm_continuous_family _ (continuous_fst.path_eval continuous_snd) + continuous_uncurry_iff.mp <| symm_continuous_family _ (continuous_fst.eval continuous_snd) @[continuity] theorem continuous_uncurry_extend_of_continuous_family {ι : Type*} [TopologicalSpace ι] diff --git a/Mathlib/Topology/ContinuousMap/Algebra.lean b/Mathlib/Topology/ContinuousMap/Algebra.lean index 4e62f17e049be..b6ff6659b4878 100644 --- a/Mathlib/Topology/ContinuousMap/Algebra.lean +++ b/Mathlib/Topology/ContinuousMap/Algebra.lean @@ -372,25 +372,27 @@ instance [CommGroup β] [TopologicalGroup β] : TopologicalGroup C(α, β) where uniformContinuous_inv.comp_tendstoUniformlyOn (tendsto_iff_forall_compact_tendstoUniformlyOn.mp Filter.tendsto_id K hK) --- TODO: rewrite the next three lemmas for products and deduce sum case via `to_additive`, once --- definition of `tprod` is in place -/-- If `α` is locally compact, and an infinite sum of functions in `C(α, β)` -converges to `g` (for the compact-open topology), then the pointwise sum converges to `g x` for -all `x ∈ α`. -/ -theorem hasSum_apply {γ : Type*} [AddCommMonoid β] [ContinuousAdd β] - {f : γ → C(α, β)} {g : C(α, β)} (hf : HasSum f g) (x : α) : - HasSum (fun i : γ => f i x) (g x) := by - let ev : C(α, β) →+ β := (Pi.evalAddMonoidHom _ x).comp coeFnAddMonoidHom - exact hf.map ev (ContinuousMap.continuous_eval_const x) - -theorem summable_apply [AddCommMonoid β] [ContinuousAdd β] {γ : Type*} {f : γ → C(α, β)} - (hf : Summable f) (x : α) : Summable fun i : γ => f i x := - (hasSum_apply hf.hasSum x).summable - -theorem tsum_apply [T2Space β] [AddCommMonoid β] [ContinuousAdd β] {γ : Type*} {f : γ → C(α, β)} - (hf : Summable f) (x : α) : - ∑' i : γ, f i x = (∑' i : γ, f i) x := - (hasSum_apply hf.hasSum x).tsum_eq +/-- If an infinite product of functions in `C(α, β)` converges to `g` +(for the compact-open topology), then the pointwise product converges to `g x` for all `x ∈ α`. -/ +@[to_additive + "If an infinite sum of functions in `C(α, β)` converges to `g` (for the compact-open topology), +then the pointwise sum converges to `g x` for all `x ∈ α`."] +theorem hasProd_apply {γ : Type*} [CommMonoid β] [ContinuousMul β] + {f : γ → C(α, β)} {g : C(α, β)} (hf : HasProd f g) (x : α) : + HasProd (fun i : γ => f i x) (g x) := by + let ev : C(α, β) →* β := (Pi.evalMonoidHom _ x).comp coeFnMonoidHom + exact hf.map ev (continuous_eval_const x) + +@[to_additive] +theorem multipliable_apply [CommMonoid β] [ContinuousMul β] {γ : Type*} {f : γ → C(α, β)} + (hf : Multipliable f) (x : α) : Multipliable fun i : γ => f i x := + (hasProd_apply hf.hasProd x).multipliable + +@[to_additive] +theorem tprod_apply [T2Space β] [CommMonoid β] [ContinuousMul β] {γ : Type*} {f : γ → C(α, β)} + (hf : Multipliable f) (x : α) : + ∏' i : γ, f i x = (∏' i : γ, f i) x := + (hasProd_apply hf.hasProd x).tprod_eq end ContinuousMap diff --git a/Mathlib/Topology/ContinuousMap/ContinuousMapZero.lean b/Mathlib/Topology/ContinuousMap/ContinuousMapZero.lean index 5fa65e8b7c342..575b4e26ad07c 100644 --- a/Mathlib/Topology/ContinuousMap/ContinuousMapZero.lean +++ b/Mathlib/Topology/ContinuousMap/ContinuousMapZero.lean @@ -85,15 +85,25 @@ lemma embedding_toContinuousMap : Embedding ((↑) : C(X, R)₀ → C(X, R)) whe inj _ _ h := ext fun x ↦ congr($(h) x) instance [T0Space R] : T0Space C(X, R)₀ := embedding_toContinuousMap.t0Space +instance [R0Space R] : R0Space C(X, R)₀ := embedding_toContinuousMap.r0Space instance [T1Space R] : T1Space C(X, R)₀ := embedding_toContinuousMap.t1Space +instance [R1Space R] : R1Space C(X, R)₀ := embedding_toContinuousMap.r1Space instance [T2Space R] : T2Space C(X, R)₀ := embedding_toContinuousMap.t2Space +instance [RegularSpace R] : RegularSpace C(X, R)₀ := embedding_toContinuousMap.regularSpace +instance [T3Space R] : T3Space C(X, R)₀ := embedding_toContinuousMap.t3Space + +instance instContinuousEvalConst : ContinuousEvalConst C(X, R)₀ X R := + .of_continuous_forget embedding_toContinuousMap.continuous + +instance instContinuousEval [LocallyCompactPair X R] : ContinuousEval C(X, R)₀ X R := + .of_continuous_forget embedding_toContinuousMap.continuous lemma closedEmbedding_toContinuousMap [T1Space R] : ClosedEmbedding ((↑) : C(X, R)₀ → C(X, R)) where toEmbedding := embedding_toContinuousMap isClosed_range := by rw [range_toContinuousMap] - exact isClosed_singleton.preimage <| ContinuousMap.continuous_eval_const 0 + exact isClosed_singleton.preimage <| continuous_eval_const 0 @[fun_prop] lemma continuous_comp_left {X Y Z : Type*} [TopologicalSpace X] diff --git a/Mathlib/Topology/Hom/ContinuousEval.lean b/Mathlib/Topology/Hom/ContinuousEval.lean new file mode 100644 index 0000000000000..7667c58b0d186 --- /dev/null +++ b/Mathlib/Topology/Hom/ContinuousEval.lean @@ -0,0 +1,69 @@ +/- +Copyright (c) 2024 Yury Kudryashov. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yury Kudryashov +-/ +import Mathlib.Topology.Hom.ContinuousEvalConst +import Mathlib.Topology.ContinuousMap.Defs + +/-! +# Bundled maps with evaluation continuous in both variables + +In this file we define a class `ContinuousEval F X Y` +saying that `F` is a bundled morphism class (in the sense of `FunLike`) +with a topology such that `fun (f, x) : F × X ↦ f x` is a continuous function. +-/ + +open scoped Topology +open Filter + +/-- A typeclass saying that `F` is a bundled morphism class (in the sense of `FunLike`) +with a topology such that `fun (f, x) : F × X ↦ f x` is a continuous function. -/ +class ContinuousEval (F : Type*) (X Y : outParam Type*) [FunLike F X Y] + [TopologicalSpace F] [TopologicalSpace X] [TopologicalSpace Y] : Prop where + /-- Evaluation of a bundled morphism at a point is continuous in both variables. -/ + continuous_eval : Continuous fun fx : F × X ↦ fx.1 fx.2 + +export ContinuousEval (continuous_eval) + +variable {F X Y Z : Type*} [FunLike F X Y] + [TopologicalSpace F] [TopologicalSpace X] [TopologicalSpace Y] [ContinuousEval F X Y] + [TopologicalSpace Z] {f : Z → F} {g : Z → X} {s : Set Z} {z : Z} + +@[continuity, fun_prop] +protected theorem Continuous.eval (hf : Continuous f) (hg : Continuous g) : + Continuous fun z ↦ f z (g z) := + continuous_eval.comp (hf.prod_mk hg) + +/-- If a type `F'` of bundled morphisms admits a continuous projection +to a type satisfying `ContinuousEval`, +then `F'` satisfies this predicate too. + +The word "forget" in the name is motivated by the term "forgetful functor". -/ +theorem ContinuousEval.of_continuous_forget {F' : Type*} [FunLike F' X Y] [TopologicalSpace F'] + {f : F' → F} (hc : Continuous f) (hf : ∀ g, ⇑(f g) = g := by intro; rfl) : + ContinuousEval F' X Y where + continuous_eval := by simpa only [← hf] using hc.fst'.eval continuous_snd + +instance (priority := 100) ContinuousEval.toContinuousMapClass : ContinuousMapClass F X Y where + map_continuous _ := continuous_const.eval continuous_id + +instance (priority := 100) ContinuousEval.toContinuousEvalConst : ContinuousEvalConst F X Y where + continuous_eval_const _ := continuous_id.eval continuous_const + +protected theorem Filter.Tendsto.eval {α : Type*} {l : Filter α} {f : α → F} {f₀ : F} + {g : α → X} {x₀ : X} (hf : Tendsto f l (𝓝 f₀)) (hg : Tendsto g l (𝓝 x₀)) : + Tendsto (fun a ↦ f a (g a)) l (𝓝 (f₀ x₀)) := + (ContinuousEval.continuous_eval.tendsto _).comp (hf.prod_mk_nhds hg) + +protected nonrec theorem ContinuousAt.eval (hf : ContinuousAt f z) (hg : ContinuousAt g z) : + ContinuousAt (fun z ↦ f z (g z)) z := + hf.eval hg + +protected nonrec theorem ContinuousWithinAt.eval (hf : ContinuousWithinAt f s z) + (hg : ContinuousWithinAt g s z) : ContinuousWithinAt (fun z ↦ f z (g z)) s z := + hf.eval hg + +protected theorem ContinuousOn.eval (hf : ContinuousOn f s) (hg : ContinuousOn g s) : + ContinuousOn (fun z ↦ f z (g z)) s := + fun z hz ↦ (hf z hz).eval (hg z hz) diff --git a/Mathlib/Topology/Hom/ContinuousEvalConst.lean b/Mathlib/Topology/Hom/ContinuousEvalConst.lean new file mode 100644 index 0000000000000..6793028a59d88 --- /dev/null +++ b/Mathlib/Topology/Hom/ContinuousEvalConst.lean @@ -0,0 +1,89 @@ +/- +Copyright (c) 2024 Yury Kudryashov. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yury Kudryashov +-/ +import Mathlib.Topology.Constructions + +/-! +# Bundled morphisms with continuous evaluation at a point + +In this file we define a typeclass +saying that `F` is a type of bundled morphisms (in the sense of `DFunLike`) +with a topology on `F` such that evaluation at a point is continuous in `f : F`. + +## Implementation Notes + +For now, we define the typeclass for non-dependent bundled functions only. +Whenever we add a type of bundled dependent functions with a topology having this property, +we may decide to generalize from `FunLike` to `DFunLike`. +-/ + +open scoped Topology +open Filter + +/-- A typeclass saying that `F` is a type of bundled morphisms (in the sense of `DFunLike`) +with a topology on `F` such that evaluation at a point is continuous in `f : F`. -/ +class ContinuousEvalConst (F : Type*) (α X : outParam Type*) [FunLike F α X] + [TopologicalSpace F] [TopologicalSpace X] : Prop where + continuous_eval_const (x : α) : Continuous fun f : F ↦ f x + +export ContinuousEvalConst (continuous_eval_const) + +section ContinuousEvalConst + +variable {F α X Z : Type*} [FunLike F α X] [TopologicalSpace F] [TopologicalSpace X] + [ContinuousEvalConst F α X] [TopologicalSpace Z] {f : Z → F} {s : Set Z} {z : Z} + +/-- If a type `F'` of bundled morphisms admits a continuous projection +to a type satisfying `ContinuousEvalConst`, +then `F'` satisfies this predicate too. + +The word "forget" in the name is motivated by the term "forgetful functor". -/ +theorem ContinuousEvalConst.of_continuous_forget {F' : Type*} [FunLike F' α X] [TopologicalSpace F'] + {f : F' → F} (hc : Continuous f) (hf : ∀ g, ⇑(f g) = g := by intro; rfl) : + ContinuousEvalConst F' α X where + continuous_eval_const x := by simpa only [← hf] using (continuous_eval_const x).comp hc + +@[continuity, fun_prop] +protected theorem Continuous.eval_const (hf : Continuous f) (x : α) : Continuous (f · x) := + (continuous_eval_const x).comp hf + +theorem continuous_coeFun : Continuous (DFunLike.coe : F → α → X) := + continuous_pi continuous_eval_const + +protected theorem Continuous.coeFun (hf : Continuous f) : Continuous fun z ↦ ⇑(f z) := + continuous_pi hf.eval_const + +protected theorem Filter.Tendsto.eval_const {ι : Type*} {l : Filter ι} {f : ι → F} {g : F} + (hf : Tendsto f l (𝓝 g)) (a : α) : Tendsto (f · a) l (𝓝 (g a)) := + ((continuous_id.eval_const a).tendsto _).comp hf + +protected theorem Filter.Tendsto.coeFun {ι : Type*} {l : Filter ι} {f : ι → F} {g : F} + (hf : Tendsto f l (𝓝 g)) : Tendsto (fun i ↦ ⇑(f i)) l (𝓝 ⇑g) := + (continuous_id.coeFun.tendsto _).comp hf + +protected nonrec theorem ContinuousAt.eval_const (hf : ContinuousAt f z) (x : α) : + ContinuousAt (f · x) z := + hf.eval_const x + +protected nonrec theorem ContinuousAt.coeFun (hf : ContinuousAt f z) : + ContinuousAt (fun z ↦ ⇑(f z)) z := + hf.coeFun + +protected nonrec theorem ContinuousWithinAt.eval_const (hf : ContinuousWithinAt f s z) (x : α) : + ContinuousWithinAt (f · x) s z := + hf.eval_const x + +protected nonrec theorem ContinuousWithinAt.coeFun (hf : ContinuousWithinAt f s z) : + ContinuousWithinAt (fun z ↦ ⇑(f z)) s z := + hf.coeFun + +protected theorem ContinuousOn.eval_const (hf : ContinuousOn f s) (x : α) : + ContinuousOn (f · x) s := + fun z hz ↦ (hf z hz).eval_const x + +protected theorem ContinuousOn.coeFun (hf : ContinuousOn f s) (x : α) : ContinuousOn (f · x) s := + fun z hz ↦ (hf z hz).eval_const x + +end ContinuousEvalConst diff --git a/Mathlib/Topology/Homotopy/HSpaces.lean b/Mathlib/Topology/Homotopy/HSpaces.lean index cf728834e55ae..fba0212ad88f5 100644 --- a/Mathlib/Topology/Homotopy/HSpaces.lean +++ b/Mathlib/Topology/Homotopy/HSpaces.lean @@ -212,7 +212,7 @@ def delayReflRight (θ : I) (γ : Path x y) : Path x y where theorem continuous_delayReflRight : Continuous fun p : I × Path x y => delayReflRight p.1 p.2 := continuous_uncurry_iff.mp <| - (continuous_snd.comp continuous_fst).path_eval <| + (continuous_snd.comp continuous_fst).eval <| continuous_qRight.comp <| continuous_snd.prod_mk <| continuous_fst.comp continuous_fst theorem delayReflRight_zero (γ : Path x y) : delayReflRight 0 γ = γ.trans (Path.refl y) := by diff --git a/Mathlib/Topology/Homotopy/HomotopyGroup.lean b/Mathlib/Topology/Homotopy/HomotopyGroup.lean index fac80ea72c606..a9ec3d6dc3be4 100644 --- a/Mathlib/Topology/Homotopy/HomotopyGroup.lean +++ b/Mathlib/Topology/Homotopy/HomotopyGroup.lean @@ -111,6 +111,11 @@ theorem ext (f g : Ω^ N X x) (H : ∀ y, f y = g y) : f = g := theorem mk_apply (f : C(I^N, X)) (H y) : (⟨f, H⟩ : Ω^ N X x) y = f y := rfl +instance instContinuousEval : ContinuousEval (Ω^ N X x) (I^N) X := + .of_continuous_forget continuous_subtype_val + +instance instContinuousEvalConst : ContinuousEvalConst (Ω^ N X x) (I^N) X := inferInstance + /-- Copy of a `GenLoop` with a new map from the unit cube equal to the old one. Useful to fix definitional equalities. -/ def copy (f : Ω^ N X x) (g : (I^N) → X) (h : g = f) : Ω^ N X x := @@ -184,7 +189,7 @@ def toLoop (i : N) (p : Ω^ N X x) : Ω (Ω^ { j // j ≠ i } X x) const where theorem continuous_toLoop (i : N) : Continuous (@toLoop N X _ x _ i) := Path.continuous_uncurry_iff.1 <| Continuous.subtype_mk - (ContinuousMap.continuous_eval.comp <| + (continuous_eval.comp <| Continuous.prodMap (ContinuousMap.continuous_curry.comp <| (ContinuousMap.continuous_comp_left _).comp continuous_subtype_val) @@ -372,7 +377,6 @@ def genLoopHomeoOfIsEmpty (N x) [IsEmpty N] : Ω^ N X x ≃ₜ X where invFun y := ⟨ContinuousMap.const _ y, fun _ ⟨i, _⟩ => isEmptyElim i⟩ left_inv f := by ext; exact congr_arg f (Subsingleton.elim _ _) right_inv _ := rfl - continuous_toFun := (ContinuousMap.continuous_eval_const (0 : N → I)).comp continuous_induced_dom continuous_invFun := ContinuousMap.const'.2.subtype_mk _ /-- The homotopy "group" indexed by an empty type is in bijection with diff --git a/Mathlib/Topology/Homotopy/Path.lean b/Mathlib/Topology/Homotopy/Path.lean index 254ed2b7f70dc..a5ad060e6071e 100644 --- a/Mathlib/Topology/Homotopy/Path.lean +++ b/Mathlib/Topology/Homotopy/Path.lean @@ -188,15 +188,7 @@ def reparam (p : Path x₀ x₁) (f : I → I) (hf : Continuous f) (hf₀ : f 0 · rw [Set.mem_singleton_iff] at hx rw [hx] simp [hf₁] - continuous_toFun := by - -- Porting note: was `continuity` in auto-param - refine continuous_const.path_eval ?_ - apply Continuous.subtype_mk - apply Continuous.add <;> apply Continuous.mul - · exact continuous_induced_dom.comp (unitInterval.continuous_symm.comp continuous_fst) - · continuity - · continuity - · continuity + continuous_toFun := by fun_prop /-- Suppose `F : Homotopy p q`. Then we have a `Homotopy p.symm q.symm` by reversing the second argument. From 8e0ebee297fc69a4ac143794e1e26ad52d25a587 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 18 Oct 2024 09:18:38 +0000 Subject: [PATCH 132/505] chore(*): fix some TODOs (#17891) Also drop an unneeded assumption in a lemma. --- .../Order/Monoid/Unbundled/ExistsOfLE.lean | 7 ++++--- Mathlib/Analysis/Calculus/ContDiff/Basic.lean | 5 ++--- Mathlib/Analysis/Oscillation.lean | 2 +- Mathlib/Data/Finset/Card.lean | 3 +-- Mathlib/Data/NNReal/Basic.lean | 20 ++++++++----------- .../Integral/RieszMarkovKakutani.lean | 2 +- .../MetricSpace/HausdorffDistance.lean | 2 +- 7 files changed, 18 insertions(+), 23 deletions(-) diff --git a/Mathlib/Algebra/Order/Monoid/Unbundled/ExistsOfLE.lean b/Mathlib/Algebra/Order/Monoid/Unbundled/ExistsOfLE.lean index da6e069dc2d22..bfca4ed38d9a9 100644 --- a/Mathlib/Algebra/Order/Monoid/Unbundled/ExistsOfLE.lean +++ b/Mathlib/Algebra/Order/Monoid/Unbundled/ExistsOfLE.lean @@ -65,20 +65,21 @@ end MulOneClass section ExistsMulOfLE variable [LinearOrder α] [DenselyOrdered α] [Monoid α] [ExistsMulOfLE α] - [CovariantClass α α (· * ·) (· < ·)] [ContravariantClass α α (· * ·) (· < ·)] {a b : α} + [ContravariantClass α α (· * ·) (· < ·)] {a b : α} @[to_additive] theorem le_of_forall_one_lt_le_mul (h : ∀ ε : α, 1 < ε → a ≤ b * ε) : a ≤ b := le_of_forall_le_of_dense fun x hxb => by obtain ⟨ε, rfl⟩ := exists_mul_of_le hxb.le - exact h _ ((lt_mul_iff_one_lt_right' b).1 hxb) + exact h _ (one_lt_of_lt_mul_right hxb) @[to_additive] theorem le_of_forall_one_lt_lt_mul' (h : ∀ ε : α, 1 < ε → a < b * ε) : a ≤ b := le_of_forall_one_lt_le_mul fun ε hε => (h ε hε).le @[to_additive] -theorem le_iff_forall_one_lt_lt_mul' : a ≤ b ↔ ∀ ε, 1 < ε → a < b * ε := +theorem le_iff_forall_one_lt_lt_mul' [CovariantClass α α (· * ·) (· < ·)] : + a ≤ b ↔ ∀ ε, 1 < ε → a < b * ε := ⟨fun h _ => lt_mul_of_le_of_one_lt h, le_of_forall_one_lt_lt_mul'⟩ end ExistsMulOfLE diff --git a/Mathlib/Analysis/Calculus/ContDiff/Basic.lean b/Mathlib/Analysis/Calculus/ContDiff/Basic.lean index f4d4b2b3d92cc..23b46aad1701c 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/Basic.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/Basic.lean @@ -1562,9 +1562,8 @@ end prodMap section AlgebraInverse -variable (𝕜) {R : Type*} [NormedRing R] --- Porting note: this couldn't be on the same line as the binder type update of `𝕜` -variable [NormedAlgebra 𝕜 R] +variable (𝕜) +variable {R : Type*} [NormedRing R] [NormedAlgebra 𝕜 R] open NormedRing ContinuousLinearMap Ring diff --git a/Mathlib/Analysis/Oscillation.lean b/Mathlib/Analysis/Oscillation.lean index c6327e3b9a059..cfc6cf9116dd9 100644 --- a/Mathlib/Analysis/Oscillation.lean +++ b/Mathlib/Analysis/Oscillation.lean @@ -52,7 +52,7 @@ namespace ContinuousWithinAt theorem oscillationWithin_eq_zero [TopologicalSpace E] {f : E → F} {D : Set E} {x : E} (hf : ContinuousWithinAt f D x) : oscillationWithin f D x = 0 := by - refine le_antisymm (le_of_forall_pos_le_add fun ε hε _ ↦ ?_) (zero_le _) + refine le_antisymm (_root_.le_of_forall_pos_le_add fun ε hε ↦ ?_) (zero_le _) rw [zero_add] have : ball (f x) (ε / 2) ∈ (𝓝[D] x).map f := hf <| ball_mem_nhds _ (by simp [ne_of_gt hε]) refine (biInf_le diam this).trans (le_of_le_of_eq diam_ball ?_) diff --git a/Mathlib/Data/Finset/Card.lean b/Mathlib/Data/Finset/Card.lean index 376e7fe70747f..67ac01fae4261 100644 --- a/Mathlib/Data/Finset/Card.lean +++ b/Mathlib/Data/Finset/Card.lean @@ -25,8 +25,7 @@ This defines the cardinality of a `Finset` and provides induction principles for -/ assert_not_exists MonoidWithZero --- TODO: After a lot more work, --- assert_not_exists OrderedCommMonoid +assert_not_exists OrderedCommMonoid open Function Multiset Nat diff --git a/Mathlib/Data/NNReal/Basic.lean b/Mathlib/Data/NNReal/Basic.lean index f0ef943d2ea06..0db61f23290b1 100644 --- a/Mathlib/Data/NNReal/Basic.lean +++ b/Mathlib/Data/NNReal/Basic.lean @@ -68,8 +68,8 @@ scoped notation "ℝ≥0" => NNReal noncomputable instance : FloorSemiring ℝ≥0 := Nonneg.floorSemiring instance instDenselyOrdered : DenselyOrdered ℝ≥0 := Nonneg.instDenselyOrdered instance : OrderBot ℝ≥0 := inferInstance -instance : Archimedean ℝ≥0 := Nonneg.instArchimedean -instance : MulArchimedean ℝ≥0 := Nonneg.instMulArchimedean +instance instArchimedean : Archimedean ℝ≥0 := Nonneg.instArchimedean +instance instMulArchimedean : MulArchimedean ℝ≥0 := Nonneg.instMulArchimedean noncomputable instance : Sub ℝ≥0 := Nonneg.sub noncomputable instance : OrderedSub ℝ≥0 := Nonneg.orderedSub @@ -493,17 +493,13 @@ theorem le_iInf_add_iInf {ι ι' : Sort*} [Nonempty ι] [Nonempty ι'] {f : ι rw [← NNReal.coe_le_coe, NNReal.coe_add, coe_iInf, coe_iInf] exact le_ciInf_add_ciInf h -example : Archimedean ℝ≥0 := by infer_instance +-- Short-circuit instance search +instance instCovariantClassAddLE : CovariantClass ℝ≥0 ℝ≥0 (· + ·) (· ≤ ·) := inferInstance +instance instContravariantClassAddLT : ContravariantClass ℝ≥0 ℝ≥0 (· + ·) (· < ·) := inferInstance +instance instCovariantClassMulLE : CovariantClass ℝ≥0 ℝ≥0 (· * ·) (· ≤ ·) := inferInstance --- Porting note (#11215): TODO: remove? -instance covariant_add : CovariantClass ℝ≥0 ℝ≥0 (· + ·) (· ≤ ·) := inferInstance - -instance contravariant_add : ContravariantClass ℝ≥0 ℝ≥0 (· + ·) (· < ·) := inferInstance - -instance covariant_mul : CovariantClass ℝ≥0 ℝ≥0 (· * ·) (· ≤ ·) := inferInstance - --- Porting note (#11215): TODO: delete? -nonrec theorem le_of_forall_pos_le_add {a b : ℝ≥0} (h : ∀ ε, 0 < ε → a ≤ b + ε) : a ≤ b := +@[deprecated le_of_forall_pos_le_add (since := "2024-10-17")] +protected theorem le_of_forall_pos_le_add {a b : ℝ≥0} (h : ∀ ε, 0 < ε → a ≤ b + ε) : a ≤ b := le_of_forall_pos_le_add h theorem lt_iff_exists_rat_btwn (a b : ℝ≥0) : diff --git a/Mathlib/MeasureTheory/Integral/RieszMarkovKakutani.lean b/Mathlib/MeasureTheory/Integral/RieszMarkovKakutani.lean index b6a9c18db7857..da3bb92699829 100644 --- a/Mathlib/MeasureTheory/Integral/RieszMarkovKakutani.lean +++ b/Mathlib/MeasureTheory/Integral/RieszMarkovKakutani.lean @@ -86,7 +86,7 @@ theorem exists_lt_rieszContentAux_add_pos (K : Compacts X) {ε : ℝ≥0} (εpos finitely subadditive: `λ(K₁ ∪ K₂) ≤ λ(K₁) + λ(K₂)` for any compact subsets `K₁, K₂ ⊆ X`. -/ theorem rieszContentAux_sup_le (K1 K2 : Compacts X) : rieszContentAux Λ (K1 ⊔ K2) ≤ rieszContentAux Λ K1 + rieszContentAux Λ K2 := by - apply NNReal.le_of_forall_pos_le_add + apply _root_.le_of_forall_pos_le_add intro ε εpos --get test functions s.t. `λ(Ki) ≤ Λfi ≤ λ(Ki) + ε/2, i=1,2` obtain ⟨f1, f_test_function_K1⟩ := exists_lt_rieszContentAux_add_pos Λ K1 (half_pos εpos) diff --git a/Mathlib/Topology/MetricSpace/HausdorffDistance.lean b/Mathlib/Topology/MetricSpace/HausdorffDistance.lean index b898627b74b4e..06799e111037e 100644 --- a/Mathlib/Topology/MetricSpace/HausdorffDistance.lean +++ b/Mathlib/Topology/MetricSpace/HausdorffDistance.lean @@ -451,7 +451,7 @@ theorem infEdist_ne_top (h : s.Nonempty) : infEdist x s ≠ ⊤ := by rcases h with ⟨y, hy⟩ exact ne_top_of_le_ne_top (edist_ne_top _ _) (infEdist_le_edist_of_mem hy) --- Porting note (#11215): TODO: make it a `simp` lemma +@[simp] theorem infEdist_eq_top_iff : infEdist x s = ∞ ↔ s = ∅ := by rcases s.eq_empty_or_nonempty with rfl | hs <;> simp [*, Nonempty.ne_empty, infEdist_ne_top] From 34f08fb6b3608c4b798f493e995f4eb615a9b01d Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 18 Oct 2024 09:52:32 +0000 Subject: [PATCH 133/505] feat: generalize results from EuclideanSpace to PiLp (#17663) The EuclideanSpace results are left in place for convenience. This generalizes leanprover-community/mathlib3#15363 Co-authored-by: Eric Wieser --- Mathlib.lean | 2 + .../Analysis/Calculus/ContDiff/WithLp.lean | 41 ++++++++++++++ Mathlib/Analysis/Calculus/FDeriv/WithLp.lean | 56 +++++++++++++++++++ .../Analysis/InnerProductSpace/Calculus.lean | 54 ++++++++---------- Mathlib/Analysis/InnerProductSpace/PiL2.lean | 13 +---- Mathlib/Analysis/Normed/Lp/PiLp.lean | 16 +++++- 6 files changed, 141 insertions(+), 41 deletions(-) create mode 100644 Mathlib/Analysis/Calculus/ContDiff/WithLp.lean create mode 100644 Mathlib/Analysis/Calculus/FDeriv/WithLp.lean diff --git a/Mathlib.lean b/Mathlib.lean index 1e5548b290acc..3c455bcceb7e9 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1004,6 +1004,7 @@ import Mathlib.Analysis.Calculus.ContDiff.Defs import Mathlib.Analysis.Calculus.ContDiff.FTaylorSeries import Mathlib.Analysis.Calculus.ContDiff.FiniteDimension import Mathlib.Analysis.Calculus.ContDiff.RCLike +import Mathlib.Analysis.Calculus.ContDiff.WithLp import Mathlib.Analysis.Calculus.Darboux import Mathlib.Analysis.Calculus.Deriv.Abs import Mathlib.Analysis.Calculus.Deriv.Add @@ -1041,6 +1042,7 @@ import Mathlib.Analysis.Calculus.FDeriv.Prod import Mathlib.Analysis.Calculus.FDeriv.RestrictScalars import Mathlib.Analysis.Calculus.FDeriv.Star import Mathlib.Analysis.Calculus.FDeriv.Symmetric +import Mathlib.Analysis.Calculus.FDeriv.WithLp import Mathlib.Analysis.Calculus.FormalMultilinearSeries import Mathlib.Analysis.Calculus.Gradient.Basic import Mathlib.Analysis.Calculus.Implicit diff --git a/Mathlib/Analysis/Calculus/ContDiff/WithLp.lean b/Mathlib/Analysis/Calculus/ContDiff/WithLp.lean new file mode 100644 index 0000000000000..4c99b51e455ba --- /dev/null +++ b/Mathlib/Analysis/Calculus/ContDiff/WithLp.lean @@ -0,0 +1,41 @@ +/- +Copyright (c) 2022 Anatole Dedecker. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Anatole Dedecker, Eric Wieser +-/ +import Mathlib.Analysis.Calculus.ContDiff.Basic +import Mathlib.Analysis.Normed.Lp.PiLp + +/-! +# Derivatives on `WithLp` +-/ + +section PiLp + +open ContinuousLinearMap + +variable {𝕜 ι : Type*} {E : ι → Type*} {H : Type*} +variable [NontriviallyNormedField 𝕜] [NormedAddCommGroup H] [∀ i, NormedAddCommGroup (E i)] + [∀ i, NormedSpace 𝕜 (E i)] [NormedSpace 𝕜 H] [Fintype ι] (p) [Fact (1 ≤ p)] + {f : H → PiLp p E} {f' : H →L[𝕜] PiLp p E} {t : Set H} {y : H} + +theorem contDiffWithinAt_piLp {n : ℕ∞} : + ContDiffWithinAt 𝕜 n f t y ↔ ∀ i, ContDiffWithinAt 𝕜 n (fun x => f x i) t y := by + rw [← (PiLp.continuousLinearEquiv p 𝕜 E).comp_contDiffWithinAt_iff, contDiffWithinAt_pi] + rfl + +theorem contDiffAt_piLp {n : ℕ∞} : + ContDiffAt 𝕜 n f y ↔ ∀ i, ContDiffAt 𝕜 n (fun x => f x i) y := by + rw [← (PiLp.continuousLinearEquiv p 𝕜 E).comp_contDiffAt_iff, contDiffAt_pi] + rfl + +theorem contDiffOn_piLp {n : ℕ∞} : + ContDiffOn 𝕜 n f t ↔ ∀ i, ContDiffOn 𝕜 n (fun x => f x i) t := by + rw [← (PiLp.continuousLinearEquiv p 𝕜 E).comp_contDiffOn_iff, contDiffOn_pi] + rfl + +theorem contDiff_piLp {n : ℕ∞} : ContDiff 𝕜 n f ↔ ∀ i, ContDiff 𝕜 n fun x => f x i := by + rw [← (PiLp.continuousLinearEquiv p 𝕜 E).comp_contDiff_iff, contDiff_pi] + rfl + +end PiLp diff --git a/Mathlib/Analysis/Calculus/FDeriv/WithLp.lean b/Mathlib/Analysis/Calculus/FDeriv/WithLp.lean new file mode 100644 index 0000000000000..1953b606870a9 --- /dev/null +++ b/Mathlib/Analysis/Calculus/FDeriv/WithLp.lean @@ -0,0 +1,56 @@ +/- +Copyright (c) 2022 Anatole Dedecker. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Anatole Dedecker, Eric Wieser +-/ +import Mathlib.Analysis.Calculus.FDeriv.Prod +import Mathlib.Analysis.Calculus.FDeriv.Equiv +import Mathlib.Analysis.Normed.Lp.PiLp + + +/-! +# Derivatives on `WithLp` +-/ + +section PiLp + +open ContinuousLinearMap + +variable {𝕜 ι : Type*} {E : ι → Type*} {H : Type*} +variable [NontriviallyNormedField 𝕜] [NormedAddCommGroup H] [∀ i, NormedAddCommGroup (E i)] + [∀ i, NormedSpace 𝕜 (E i)] [NormedSpace 𝕜 H] [Fintype ι] (p) [Fact (1 ≤ p)] + {f : H → PiLp p E} {f' : H →L[𝕜] PiLp p E} {t : Set H} {y : H} + +theorem differentiableWithinAt_piLp : + DifferentiableWithinAt 𝕜 f t y ↔ ∀ i, DifferentiableWithinAt 𝕜 (fun x => f x i) t y := by + rw [← (PiLp.continuousLinearEquiv p 𝕜 E).comp_differentiableWithinAt_iff, + differentiableWithinAt_pi] + rfl + +theorem differentiableAt_piLp : + DifferentiableAt 𝕜 f y ↔ ∀ i, DifferentiableAt 𝕜 (fun x => f x i) y := by + rw [← (PiLp.continuousLinearEquiv p 𝕜 E).comp_differentiableAt_iff, differentiableAt_pi] + rfl + +theorem differentiableOn_piLp : + DifferentiableOn 𝕜 f t ↔ ∀ i, DifferentiableOn 𝕜 (fun x => f x i) t := by + rw [← (PiLp.continuousLinearEquiv p 𝕜 E).comp_differentiableOn_iff, differentiableOn_pi] + rfl + +theorem differentiable_piLp : Differentiable 𝕜 f ↔ ∀ i, Differentiable 𝕜 fun x => f x i := by + rw [← (PiLp.continuousLinearEquiv p 𝕜 E).comp_differentiable_iff, differentiable_pi] + rfl + +theorem hasStrictFDerivAt_piLp : + HasStrictFDerivAt f f' y ↔ + ∀ i, HasStrictFDerivAt (fun x => f x i) (PiLp.proj _ _ i ∘L f') y := by + rw [← (PiLp.continuousLinearEquiv p 𝕜 E).comp_hasStrictFDerivAt_iff, hasStrictFDerivAt_pi'] + rfl + +theorem hasFDerivWithinAt_piLp : + HasFDerivWithinAt f f' t y ↔ + ∀ i, HasFDerivWithinAt (fun x => f x i) (PiLp.proj _ _ i ∘L f') t y := by + rw [← (PiLp.continuousLinearEquiv p 𝕜 E).comp_hasFDerivWithinAt_iff, hasFDerivWithinAt_pi'] + rfl + +end PiLp diff --git a/Mathlib/Analysis/InnerProductSpace/Calculus.lean b/Mathlib/Analysis/InnerProductSpace/Calculus.lean index 44808e4d72423..018aa26547ec5 100644 --- a/Mathlib/Analysis/InnerProductSpace/Calculus.lean +++ b/Mathlib/Analysis/InnerProductSpace/Calculus.lean @@ -6,6 +6,8 @@ Authors: Yury Kudryashov import Mathlib.Analysis.InnerProductSpace.PiL2 import Mathlib.Analysis.SpecialFunctions.Sqrt import Mathlib.Analysis.NormedSpace.HomeomorphBall +import Mathlib.Analysis.Calculus.ContDiff.WithLp +import Mathlib.Analysis.Calculus.FDeriv.WithLp /-! # Calculus in inner product spaces @@ -264,60 +266,52 @@ end DerivInner section PiLike +/-! ### Convenience aliases of `PiLp` lemmas for `EuclideanSpace` -/ + open ContinuousLinearMap variable {𝕜 ι H : Type*} [RCLike 𝕜] [NormedAddCommGroup H] [NormedSpace 𝕜 H] [Fintype ι] {f : H → EuclideanSpace 𝕜 ι} {f' : H →L[𝕜] EuclideanSpace 𝕜 ι} {t : Set H} {y : H} theorem differentiableWithinAt_euclidean : - DifferentiableWithinAt 𝕜 f t y ↔ ∀ i, DifferentiableWithinAt 𝕜 (fun x => f x i) t y := by - rw [← (EuclideanSpace.equiv ι 𝕜).comp_differentiableWithinAt_iff, differentiableWithinAt_pi] - rfl + DifferentiableWithinAt 𝕜 f t y ↔ ∀ i, DifferentiableWithinAt 𝕜 (fun x => f x i) t y := + differentiableWithinAt_piLp _ theorem differentiableAt_euclidean : - DifferentiableAt 𝕜 f y ↔ ∀ i, DifferentiableAt 𝕜 (fun x => f x i) y := by - rw [← (EuclideanSpace.equiv ι 𝕜).comp_differentiableAt_iff, differentiableAt_pi] - rfl + DifferentiableAt 𝕜 f y ↔ ∀ i, DifferentiableAt 𝕜 (fun x => f x i) y := + differentiableAt_piLp _ theorem differentiableOn_euclidean : - DifferentiableOn 𝕜 f t ↔ ∀ i, DifferentiableOn 𝕜 (fun x => f x i) t := by - rw [← (EuclideanSpace.equiv ι 𝕜).comp_differentiableOn_iff, differentiableOn_pi] - rfl + DifferentiableOn 𝕜 f t ↔ ∀ i, DifferentiableOn 𝕜 (fun x => f x i) t := + differentiableOn_piLp _ -theorem differentiable_euclidean : Differentiable 𝕜 f ↔ ∀ i, Differentiable 𝕜 fun x => f x i := by - rw [← (EuclideanSpace.equiv ι 𝕜).comp_differentiable_iff, differentiable_pi] - rfl +theorem differentiable_euclidean : Differentiable 𝕜 f ↔ ∀ i, Differentiable 𝕜 fun x => f x i := + differentiable_piLp _ theorem hasStrictFDerivAt_euclidean : HasStrictFDerivAt f f' y ↔ - ∀ i, HasStrictFDerivAt (fun x => f x i) (EuclideanSpace.proj i ∘L f') y := by - rw [← (EuclideanSpace.equiv ι 𝕜).comp_hasStrictFDerivAt_iff, hasStrictFDerivAt_pi'] - rfl + ∀ i, HasStrictFDerivAt (fun x => f x i) (PiLp.proj _ _ i ∘L f') y := + hasStrictFDerivAt_piLp _ theorem hasFDerivWithinAt_euclidean : HasFDerivWithinAt f f' t y ↔ - ∀ i, HasFDerivWithinAt (fun x => f x i) (EuclideanSpace.proj i ∘L f') t y := by - rw [← (EuclideanSpace.equiv ι 𝕜).comp_hasFDerivWithinAt_iff, hasFDerivWithinAt_pi'] - rfl + ∀ i, HasFDerivWithinAt (fun x => f x i) (PiLp.proj _ _ i ∘L f') t y := + hasFDerivWithinAt_piLp _ theorem contDiffWithinAt_euclidean {n : ℕ∞} : - ContDiffWithinAt 𝕜 n f t y ↔ ∀ i, ContDiffWithinAt 𝕜 n (fun x => f x i) t y := by - rw [← (EuclideanSpace.equiv ι 𝕜).comp_contDiffWithinAt_iff, contDiffWithinAt_pi] - rfl + ContDiffWithinAt 𝕜 n f t y ↔ ∀ i, ContDiffWithinAt 𝕜 n (fun x => f x i) t y := + contDiffWithinAt_piLp _ theorem contDiffAt_euclidean {n : ℕ∞} : - ContDiffAt 𝕜 n f y ↔ ∀ i, ContDiffAt 𝕜 n (fun x => f x i) y := by - rw [← (EuclideanSpace.equiv ι 𝕜).comp_contDiffAt_iff, contDiffAt_pi] - rfl + ContDiffAt 𝕜 n f y ↔ ∀ i, ContDiffAt 𝕜 n (fun x => f x i) y := + contDiffAt_piLp _ theorem contDiffOn_euclidean {n : ℕ∞} : - ContDiffOn 𝕜 n f t ↔ ∀ i, ContDiffOn 𝕜 n (fun x => f x i) t := by - rw [← (EuclideanSpace.equiv ι 𝕜).comp_contDiffOn_iff, contDiffOn_pi] - rfl + ContDiffOn 𝕜 n f t ↔ ∀ i, ContDiffOn 𝕜 n (fun x => f x i) t := + contDiffOn_piLp _ -theorem contDiff_euclidean {n : ℕ∞} : ContDiff 𝕜 n f ↔ ∀ i, ContDiff 𝕜 n fun x => f x i := by - rw [← (EuclideanSpace.equiv ι 𝕜).comp_contDiff_iff, contDiff_pi] - rfl +theorem contDiff_euclidean {n : ℕ∞} : ContDiff 𝕜 n f ↔ ∀ i, ContDiff 𝕜 n fun x => f x i := + contDiff_piLp _ end PiLike diff --git a/Mathlib/Analysis/InnerProductSpace/PiL2.lean b/Mathlib/Analysis/InnerProductSpace/PiL2.lean index 2c14b91ede9f1..3a77d1f7c5328 100644 --- a/Mathlib/Analysis/InnerProductSpace/PiL2.lean +++ b/Mathlib/Analysis/InnerProductSpace/PiL2.lean @@ -200,18 +200,11 @@ abbrev EuclideanSpace.equiv : EuclideanSpace 𝕜 ι ≃L[𝕜] ι → 𝕜 := variable {ι 𝕜} --- TODO : This should be generalized to `PiLp`. /-- The projection on the `i`-th coordinate of `EuclideanSpace 𝕜 ι`, as a linear map. -/ -@[simps!] -def EuclideanSpace.projₗ (i : ι) : EuclideanSpace 𝕜 ι →ₗ[𝕜] 𝕜 := - (LinearMap.proj i).comp (WithLp.linearEquiv 2 𝕜 (ι → 𝕜) : EuclideanSpace 𝕜 ι →ₗ[𝕜] ι → 𝕜) +abbrev EuclideanSpace.projₗ (i : ι) : EuclideanSpace 𝕜 ι →ₗ[𝕜] 𝕜 := PiLp.projₗ _ _ i --- TODO : This should be generalized to `PiLp`. -/-- The projection on the `i`-th coordinate of `EuclideanSpace 𝕜 ι`, -as a continuous linear map. -/ -@[simps! apply coe] -def EuclideanSpace.proj (i : ι) : EuclideanSpace 𝕜 ι →L[𝕜] 𝕜 := - ⟨EuclideanSpace.projₗ i, continuous_apply i⟩ +/-- The projection on the `i`-th coordinate of `EuclideanSpace 𝕜 ι`, as a continuous linear map. -/ +abbrev EuclideanSpace.proj (i : ι) : EuclideanSpace 𝕜 ι →L[𝕜] 𝕜 := PiLp.proj _ _ i section DecEq diff --git a/Mathlib/Analysis/Normed/Lp/PiLp.lean b/Mathlib/Analysis/Normed/Lp/PiLp.lean index 3902519c5b489..3b9cc58963b91 100644 --- a/Mathlib/Analysis/Normed/Lp/PiLp.lean +++ b/Mathlib/Analysis/Normed/Lp/PiLp.lean @@ -91,7 +91,7 @@ section /- Register simplification lemmas for the applications of `PiLp` elements, as the usual lemmas for Pi types will not trigger. -/ variable {𝕜 p α} -variable [SeminormedRing 𝕜] [∀ i, SeminormedAddCommGroup (β i)] +variable [Semiring 𝕜] [∀ i, SeminormedAddCommGroup (β i)] variable [∀ i, Module 𝕜 (β i)] (c : 𝕜) variable (x y : PiLp p β) (i : ι) @@ -121,6 +121,13 @@ theorem smul_apply : (c • x) i = c • x i := @[simp] theorem neg_apply : (-x) i = -x i := rfl + +variable (p) in +/-- The projection on the `i`-th coordinate of `WithLp p (∀ i, α i)`, as a linear map. -/ +@[simps!] +def projₗ (i : ι) : PiLp p β →ₗ[𝕜] β i := + (LinearMap.proj i : (∀ i, β i) →ₗ[𝕜] β i) ∘ₗ (WithLp.linearEquiv p 𝕜 (∀ i, β i)).toLinearMap + end /-! Note that the unapplied versions of these lemmas are deliberately omitted, as they break @@ -891,6 +898,13 @@ protected def continuousLinearEquiv : PiLp p β ≃L[𝕜] ∀ i, β i where continuous_toFun := continuous_equiv _ _ continuous_invFun := continuous_equiv_symm _ _ +variable {𝕜} in +/-- The projection on the `i`-th coordinate of `PiLp p β`, as a continuous linear map. -/ +@[simps!] +def proj (i : ι) : PiLp p β →L[𝕜] β i where + __ := projₗ p β i + cont := continuous_apply i + end Fintype section Basis From 6338c7313383353c208346451dbdf0cc037ac8a2 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 18 Oct 2024 09:52:33 +0000 Subject: [PATCH 134/505] fix: improve line-wrapping of (D)Finsupp reprs (#17882) These were previously wrapping very poorly. Also update tests to use `#guard_msgs` more effectively, and fix a typo in the delaborators for DFinsupp. --- Mathlib/Data/DFinsupp/Notation.lean | 10 +++---- Mathlib/Data/Finsupp/Notation.lean | 6 ++-- test/dfinsupp_notation.lean | 41 ++++++++++++++++++++------- test/finsupp_notation.lean | 44 +++++++++++++++++++++++++---- 4 files changed, 78 insertions(+), 23 deletions(-) diff --git a/Mathlib/Data/DFinsupp/Notation.lean b/Mathlib/Data/DFinsupp/Notation.lean index f4782385250a8..bdb9cdad1a53f 100644 --- a/Mathlib/Data/DFinsupp/Notation.lean +++ b/Mathlib/Data/DFinsupp/Notation.lean @@ -44,13 +44,13 @@ def elabUpdate₀ : Elab.Term.TermElab | _ => fun _ => Elab.throwUnsupportedSyntax /-- Unexpander for the `fun₀ | i => x` notation. -/ -@[app_unexpander Finsupp.single] +@[app_unexpander DFinsupp.single] def singleUnexpander : Lean.PrettyPrinter.Unexpander | `($_ $pat $val) => `(fun₀ | $pat => $val) | _ => throw () /-- Unexpander for the `fun₀ | i => x` notation. -/ -@[app_unexpander Finsupp.update] +@[app_unexpander DFinsupp.update] def updateUnexpander : Lean.PrettyPrinter.Unexpander | `($_ $f $pat $val) => match f with | `(fun₀ $xs:matchAlt*) => `(fun₀ $xs:matchAlt* | $pat => $val) @@ -68,9 +68,9 @@ unsafe instance {α : Type*} {β : α → Type*} [Repr α] [∀ i, Repr (β i)] if vals.length = 0 then "0" else - let ret := "fun₀" ++ - Std.Format.join (vals_dedup.map <| - fun a => f!" | " ++ a.1 ++ f!" => " ++ a.2) + let ret : Std.Format := f!"fun₀" ++ .nest 2 ( + .group (.join <| vals_dedup.map fun a => + .line ++ .group (f!"| {a.1} =>" ++ .line ++ a.2))) if p ≥ leadPrec then Format.paren ret else ret end DFinsupp diff --git a/Mathlib/Data/Finsupp/Notation.lean b/Mathlib/Data/Finsupp/Notation.lean index 2091f9b3cb8d1..12c11f9e628e5 100644 --- a/Mathlib/Data/Finsupp/Notation.lean +++ b/Mathlib/Data/Finsupp/Notation.lean @@ -83,9 +83,9 @@ unsafe instance instRepr {α β} [Repr α] [Repr β] [Zero β] : Repr (α →₀ if f.support.card = 0 then "0" else - let ret := "fun₀" ++ - Std.Format.join (f.support.val.unquot.map <| - fun a => " | " ++ repr a ++ " => " ++ repr (f a)) + let ret : Std.Format := f!"fun₀" ++ .nest 2 ( + .group (.join <| f.support.val.unquot.map fun a => + .line ++ .group (f!"| {repr a} =>" ++ .line ++ repr (f a)))) if p ≥ leadPrec then Format.paren ret else ret -- This cannot be put in `Mathlib.Data.DFinsupp.Notation` where it belongs, since doc-strings diff --git a/test/dfinsupp_notation.lean b/test/dfinsupp_notation.lean index ed6511532b5f9..ea4d3815079ae 100644 --- a/test/dfinsupp_notation.lean +++ b/test/dfinsupp_notation.lean @@ -7,20 +7,41 @@ example : (fun₀ | 1 | 2 | 3 => 3 | 3 => 4 : Π₀ i, Fin (i + 10)) 1 = 3 := by example : (fun₀ | 1 | 2 | 3 => 3 | 3 => 4 : Π₀ i, Fin (i + 10)) 2 = 3 := by simp example : (fun₀ | 1 | 2 | 3 => 3 | 3 => 4 : Π₀ i, Fin (i + 10)) 3 = 4 := by simp +section Repr + +/-- info: fun₀ | 1 => 3 | 2 => 3 : Π₀ (i : ℕ), Fin (i + 10) -/ +#guard_msgs in +#check (fun₀ | 1 => 3 | 2 => 3 : Π₀ i, Fin (i + 10)) + +/-- info: fun₀ | 2 => 7 -/ +#guard_msgs in +#eval ((fun₀ | 1 => 3 | 2 => 3) + (fun₀ | 1 => -3 | 2 => 4) : Π₀ _, ℤ) + /-- -info: +info: fun₀ + | ["there are five words here", "and five more words here"] => 5 + | ["there are seven words but only here"] => 7 + | ["just two"] => 2 -/ #guard_msgs in -#eval show Lean.MetaM Unit from - guard <| - reprStr (fun₀ | 1 => 3 | 2 => 3 : Π₀ i, Fin (i + 10)) - = "fun₀ | 1 => 3 | 2 => 3" +#eval (fun₀ + | ["there are five words here", "and five more words here"] => 5 + | ["there are seven words but only here"] => 7 + | ["just two"] => 2 : Π₀ _ : List String, ℕ) + +end Repr +section PrettyPrinter /-- -info: +info: fun₀ + | ["there are five words here", "and five more words here"] => 5 + | ["there are seven words but only here"] => 7 + | ["just two"] => 2 : Π₀ (i : List String), ℕ -/ #guard_msgs in -#eval show Lean.MetaM Unit from - guard <| - reprStr ((fun₀ | 1 => 3 | 2 => 3) + (fun₀ | 1 => -3 | 2 => 4) : Π₀ _, ℤ) - = "fun₀ | 2 => 7" +#check (fun₀ + | ["there are five words here", "and five more words here"] => 5 + | ["there are seven words but only here"] => 7 + | ["just two"] => 2 : Π₀ _ : List String, ℕ) + +end PrettyPrinter diff --git a/test/finsupp_notation.lean b/test/finsupp_notation.lean index e042b67a8903f..08e871dde99cb 100644 --- a/test/finsupp_notation.lean +++ b/test/finsupp_notation.lean @@ -12,14 +12,48 @@ example : (fun₀ | 1 | 2 | 3 => 3 | 3 => 4) 2 = 3 := by example : (fun₀ | 1 | 2 | 3 => 3 | 3 => 4) 3 = 4 := by simp +section repr + +/-- info: fun₀ | 1 => 3 | 2 => 3 -/ +#guard_msgs in +#eval (Finsupp.mk {1, 2} (fun | 1 | 2 => 3 | _ => 0) (fun x => by aesop)) + /-- -info: +info: fun₀ + | ["there are five words here", "and five more words here"] => 5 + | ["there are seven words but only here"] => 7 + | ["just two"] => 2 -/ #guard_msgs in -#eval show Lean.MetaM Unit from - guard <| - reprStr (Finsupp.mk {1, 2} (fun | 1 | 2 => 3 | _ => 0) (fun x => by aesop)) - = "fun₀ | 1 => 3 | 2 => 3" +#eval Finsupp.mk + {["there are five words here", "and five more words here"], + ["there are seven words but only here"], + ["just two"]} + (fun + | ["there are five words here", "and five more words here"] => 5 + | ["there are seven words but only here"] => 7 + | ["just two"] => 2 + | _ => 0) + (fun x => by aesop) + +end repr + +section PrettyPrinter + +/-- +info: fun₀ + | ["there are five words here", "and five more words here"] => 5 + | ["there are seven words but only here"] => 7 + | ["just two"] => 2 : List String →₀ ℕ +-/ +#guard_msgs in +#check fun₀ + | ["there are five words here", "and five more words here"] => 5 + | ["there are seven words but only here"] => 7 + | ["just two"] => 2 + +end PrettyPrinter + /-! ## (computable) number theory examples -/ From fd77f31972aa69d64ea2889b8d154417042e4aba Mon Sep 17 00:00:00 2001 From: Jakob von Raumer Date: Fri, 18 Oct 2024 09:52:34 +0000 Subject: [PATCH 135/505] chore(CategoryTheory): Make terminology for final functors consistent (#17895) Removes the remaining mentions of "cofinal functors". --- Mathlib/CategoryTheory/Limits/Final.lean | 32 ++++++++++----------- Mathlib/CategoryTheory/Limits/Presheaf.lean | 2 +- 2 files changed, 17 insertions(+), 17 deletions(-) diff --git a/Mathlib/CategoryTheory/Limits/Final.lean b/Mathlib/CategoryTheory/Limits/Final.lean index 2341721a15f55..965ee9e0e669e 100644 --- a/Mathlib/CategoryTheory/Limits/Final.lean +++ b/Mathlib/CategoryTheory/Limits/Final.lean @@ -36,7 +36,7 @@ This readily implies 2., as `comp_hasColimit`, `hasColimit_of_comp`, and `colimi From 2. we can specialize to `G = coyoneda.obj (op d)` to obtain 3., as `colimitCompCoyonedaIso`. -From 3., we prove 1. directly in `cofinal_of_colimit_comp_coyoneda_iso_pUnit`. +From 3., we prove 1. directly in `final_of_colimit_comp_coyoneda_iso_pUnit`. Dually, we prove that if a functor `F : C ⥤ D` is initial, then any functor `G : D ⥤ E` has a limit if and only if `F ⋙ G` does, and these limits are isomorphic via `limit.pre G F`. @@ -170,13 +170,13 @@ instance (d : D) : Nonempty (StructuredArrow d F) := variable {E : Type u₃} [Category.{v₃} E] (G : D ⥤ E) /-- -When `F : C ⥤ D` is cofinal, we denote by `lift F d` an arbitrary choice of object in `C` such that +When `F : C ⥤ D` is final, we denote by `lift F d` an arbitrary choice of object in `C` such that there exists a morphism `d ⟶ F.obj (lift F d)`. -/ def lift (d : D) : C := (Classical.arbitrary (StructuredArrow d F)).right -/-- When `F : C ⥤ D` is cofinal, we denote by `homToLift` an arbitrary choice of morphism +/-- When `F : C ⥤ D` is final, we denote by `homToLift` an arbitrary choice of morphism `d ⟶ F.obj (lift F d)`. -/ def homToLift (d : D) : d ⟶ F.obj (lift F d) := @@ -267,7 +267,7 @@ theorem colimit_cocone_comp_aux (s : Cocone (F ⋙ G)) (j : C) : variable (F G) -/-- If `F` is cofinal, +/-- If `F` is final, the category of cocones on `F ⋙ G` is equivalent to the category of cocones on `G`, for any `G : D ⥤ E`. -/ @@ -280,13 +280,13 @@ def coconesEquiv : Cocone (F ⋙ G) ≌ Cocone G where variable {G} -/-- When `F : C ⥤ D` is cofinal, and `t : Cocone G` for some `G : D ⥤ E`, +/-- When `F : C ⥤ D` is final, and `t : Cocone G` for some `G : D ⥤ E`, `t.whisker F` is a colimit cocone exactly when `t` is. -/ def isColimitWhiskerEquiv (t : Cocone G) : IsColimit (t.whisker F) ≃ IsColimit t := IsColimit.ofCoconeEquiv (coconesEquiv F G).symm -/-- When `F` is cofinal, and `t : Cocone (F ⋙ G)`, +/-- When `F` is final, and `t : Cocone (F ⋙ G)`, `extendCocone.obj t` is a colimit cocone exactly when `t` is. -/ def isColimitExtendCoconeEquiv (t : Cocone (F ⋙ G)) : @@ -312,7 +312,7 @@ section variable (G) -/-- When `F : C ⥤ D` is cofinal, and `G : D ⥤ E` has a colimit, then `F ⋙ G` has a colimit also and +/-- When `F : C ⥤ D` is final, and `G : D ⥤ E` has a colimit, then `F ⋙ G` has a colimit also and `colimit (F ⋙ G) ≅ colimit G` https://stacks.math.columbia.edu/tag/04E7 @@ -328,7 +328,7 @@ def colimitCoconeOfComp (t : ColimitCocone (F ⋙ G)) : ColimitCocone G where cocone := extendCocone.obj t.cocone isColimit := (isColimitExtendCoconeEquiv F _).symm t.isColimit -/-- When `F` is cofinal, and `F ⋙ G` has a colimit, then `G` has a colimit also. +/-- When `F` is final, and `F ⋙ G` has a colimit, then `G` has a colimit also. We can't make this an instance, because `F` is not determined by the goal. (Even if this weren't a problem, it would cause a loop with `comp_hasColimit`.) @@ -345,7 +345,7 @@ section -- Porting note: this instance does not seem to be found automatically --attribute [local instance] hasColimit_of_comp -/-- When `F` is cofinal, and `F ⋙ G` has a colimit, then `G` has a colimit also and +/-- When `F` is final, and `F ⋙ G` has a colimit, then `G` has a colimit also and `colimit (F ⋙ G) ≅ colimit G` https://stacks.math.columbia.edu/tag/04E7 @@ -390,9 +390,9 @@ theorem zigzag_of_eqvGen_quot_rel {F : C ⥤ D} {d : D} {f₁ f₂ : ΣX, d ⟶ end Final -/-- If `colimit (F ⋙ coyoneda.obj (op d)) ≅ PUnit` for all `d : D`, then `F` is cofinal. +/-- If `colimit (F ⋙ coyoneda.obj (op d)) ≅ PUnit` for all `d : D`, then `F` is final. -/ -theorem cofinal_of_colimit_comp_coyoneda_iso_pUnit +theorem final_of_colimit_comp_coyoneda_iso_pUnit (I : ∀ d, colimit (F ⋙ coyoneda.obj (op d)) ≅ PUnit) : Final F := ⟨fun d => by have : Nonempty (StructuredArrow d F) := by @@ -410,18 +410,18 @@ theorem cofinal_of_colimit_comp_coyoneda_iso_pUnit clear e y₁ y₂ exact Final.zigzag_of_eqvGen_quot_rel t⟩ -/-- A variant of `cofinal_of_colimit_comp_coyoneda_iso_pUnit` where we bind the various claims +/-- A variant of `final_of_colimit_comp_coyoneda_iso_pUnit` where we bind the various claims about `colimit (F ⋙ coyoneda.obj (Opposite.op d))` for each `d : D` into a single claim about the presheaf `colimit (F ⋙ yoneda)`. -/ -theorem cofinal_of_isTerminal_colimit_comp_yoneda +theorem final_of_isTerminal_colimit_comp_yoneda (h : IsTerminal (colimit (F ⋙ yoneda))) : Final F := by - refine cofinal_of_colimit_comp_coyoneda_iso_pUnit _ (fun d => ?_) + refine final_of_colimit_comp_coyoneda_iso_pUnit _ (fun d => ?_) refine Types.isTerminalEquivIsoPUnit _ ?_ let b := IsTerminal.isTerminalObj ((evaluation _ _).obj (Opposite.op d)) _ h exact b.ofIso <| preservesColimitIso ((evaluation _ _).obj (Opposite.op d)) (F ⋙ yoneda) /-- If the universal morphism `colimit (F ⋙ coyoneda.obj (op d)) ⟶ colimit (coyoneda.obj (op d))` -is an isomorphism (as it always is when `F` is cofinal), +is an isomorphism (as it always is when `F` is final), then `colimit (F ⋙ coyoneda.obj (op d)) ≅ PUnit` (simply because `colimit (coyoneda.obj (op d)) ≅ PUnit`). -/ @@ -437,7 +437,7 @@ variable {C : Type v} [Category.{v} C] {D : Type v} [Category.{v} D] (F : C ⥤ theorem final_iff_isIso_colimit_pre : Final F ↔ ∀ G : D ⥤ Type v, IsIso (colimit.pre G F) := ⟨fun _ => inferInstance, - fun _ => cofinal_of_colimit_comp_coyoneda_iso_pUnit _ fun _ => Final.colimitCompCoyonedaIso _ _⟩ + fun _ => final_of_colimit_comp_coyoneda_iso_pUnit _ fun _ => Final.colimitCompCoyonedaIso _ _⟩ end SmallCategory diff --git a/Mathlib/CategoryTheory/Limits/Presheaf.lean b/Mathlib/CategoryTheory/Limits/Presheaf.lean index 5dc588dc02b38..79b189b11e50c 100644 --- a/Mathlib/CategoryTheory/Limits/Presheaf.lean +++ b/Mathlib/CategoryTheory/Limits/Presheaf.lean @@ -559,7 +559,7 @@ variable {I : Type v₁} [SmallCategory I] (F : I ⥤ C) Proposition 2.6.3(ii) in [Kashiwara2006] -/ theorem final_toCostructuredArrow_comp_pre {c : Cocone (F ⋙ yoneda)} (hc : IsColimit c) : Functor.Final (c.toCostructuredArrow ⋙ CostructuredArrow.pre F yoneda c.pt) := by - apply Functor.cofinal_of_isTerminal_colimit_comp_yoneda + apply Functor.final_of_isTerminal_colimit_comp_yoneda suffices IsTerminal (colimit ((c.toCostructuredArrow ⋙ CostructuredArrow.pre F yoneda c.pt) ⋙ CostructuredArrow.toOver yoneda c.pt)) by From 23752ee05683cc476abb6f9eb98b63c0eb5d3da1 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Fri, 18 Oct 2024 10:20:47 +0000 Subject: [PATCH 136/505] feat: the algebraic closure of an algebraic k-algebra is an algebraic closure of k (#17894) Co-authored-by: FR --- Mathlib/FieldTheory/IsAlgClosed/AlgebraicClosure.lean | 5 +++++ 1 file changed, 5 insertions(+) diff --git a/Mathlib/FieldTheory/IsAlgClosed/AlgebraicClosure.lean b/Mathlib/FieldTheory/IsAlgClosed/AlgebraicClosure.lean index ea7fe25e09f58..d58f8dcbc128b 100644 --- a/Mathlib/FieldTheory/IsAlgClosed/AlgebraicClosure.lean +++ b/Mathlib/FieldTheory/IsAlgClosed/AlgebraicClosure.lean @@ -403,4 +403,9 @@ instance [CharZero k] : CharZero (AlgebraicClosure k) := instance {p : ℕ} [CharP k p] : CharP (AlgebraicClosure k) p := charP_of_injective_algebraMap (RingHom.injective (algebraMap k (AlgebraicClosure k))) p +instance {L : Type*} [Field k] [Field L] [Algebra k L] [Algebra.IsAlgebraic k L] : + IsAlgClosure k (AlgebraicClosure L) where + isAlgebraic := .trans (L := L) + isAlgClosed := inferInstance + end AlgebraicClosure From 1cc3f60d578ae8b32e7e308018e9f409a5dbaedc Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Fri, 18 Oct 2024 10:35:29 +0000 Subject: [PATCH 137/505] chore(Order/RelIso/Set): golf `wellFounded_iff_wellFounded_subrel` (#17606) We also move it out of the initial segment file - it doesn't need that machinery to be proven. --- Mathlib/Order/InitialSeg.lean | 14 -------------- Mathlib/Order/RelIso/Set.lean | 11 +++++++++++ 2 files changed, 11 insertions(+), 14 deletions(-) diff --git a/Mathlib/Order/InitialSeg.lean b/Mathlib/Order/InitialSeg.lean index 513501c25072e..e53d78efa0277 100644 --- a/Mathlib/Order/InitialSeg.lean +++ b/Mathlib/Order/InitialSeg.lean @@ -461,20 +461,6 @@ protected theorem acc [IsTrans β s] (f : r ≺i s) (a : α) : Acc r a ↔ Acc s end PrincipalSeg -/-- A relation is well-founded iff every principal segment of it is well-founded. - -In this lemma we use `Subrel` to indicate its principal segments because it's usually more -convenient to use. --/ -theorem wellFounded_iff_wellFounded_subrel {β : Type*} {s : β → β → Prop} [IsTrans β s] : - WellFounded s ↔ ∀ b, WellFounded (Subrel s { b' | s b' b }) := by - refine - ⟨fun wf b => ⟨fun b' => ((PrincipalSeg.ofElement _ b).acc b').mpr (wf.apply b')⟩, fun wf => - ⟨fun b => Acc.intro _ fun b' hb' => ?_⟩⟩ - let f := PrincipalSeg.ofElement s b - obtain ⟨b', rfl⟩ := f.mem_range_of_rel_top ((PrincipalSeg.ofElement_top s b).symm ▸ hb') - exact (f.acc b').mp ((wf b).apply b') - theorem wellFounded_iff_principalSeg.{u} {β : Type u} {s : β → β → Prop} [IsTrans β s] : WellFounded s ↔ ∀ (α : Type u) (r : α → α → Prop) (_ : r ≺i s), WellFounded r := ⟨fun wf _ _ f => RelHomClass.wellFounded f.toRelEmbedding wf, fun h => diff --git a/Mathlib/Order/RelIso/Set.lean b/Mathlib/Order/RelIso/Set.lean index 961b9773f5912..ec7b62b18b16f 100644 --- a/Mathlib/Order/RelIso/Set.lean +++ b/Mathlib/Order/RelIso/Set.lean @@ -98,3 +98,14 @@ theorem RelIso.preimage_eq_image_symm (e : r ≃r s) (t : Set β) : e ⁻¹' t = rw [e.symm.image_eq_preimage_symm]; rfl end image + +theorem Acc.of_subrel {r : α → α → Prop} [IsTrans α r] {b : α} (a : { a // r a b }) + (h : Acc (Subrel r { a | r a b }) a) : Acc r a.1 := + h.recOn fun a _ IH ↦ ⟨_, fun _ hb ↦ IH ⟨_, _root_.trans hb a.2⟩ hb⟩ + +/-- A relation `r` is well-founded iff every downward-interval `{ a | r a b }` of it is +well-founded. -/ +theorem wellFounded_iff_wellFounded_subrel {r : α → α → Prop} [IsTrans α r] : + WellFounded r ↔ ∀ b, WellFounded (Subrel r { a | r a b }) where + mp h _ := InvImage.wf Subtype.val h + mpr h := ⟨fun a ↦ ⟨_, fun b hr ↦ ((h a).apply _).of_subrel ⟨b, hr⟩⟩⟩ From cbaeaf5ebf33c9f777f35f6cf3a44029e7d35c0e Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 18 Oct 2024 10:35:30 +0000 Subject: [PATCH 138/505] chore(Topology): deprecate `Homeomorph.toContinuousMap` (#17714) Use `_root_.toContinuousMap` instead. --- Mathlib/MeasureTheory/Measure/Content.lean | 6 +++--- Mathlib/Topology/Category/TopCat/Basic.lean | 4 ++-- Mathlib/Topology/ContinuousMap/Algebra.lean | 4 ++-- Mathlib/Topology/ContinuousMap/Basic.lean | 16 ++++++---------- Mathlib/Topology/ContinuousMap/Compact.lean | 4 ++-- .../ContinuousMap/ContinuousMapZero.lean | 8 ++++---- Mathlib/Topology/ContinuousMap/Defs.lean | 4 ++++ .../Topology/ContinuousMap/Polynomial.lean | 2 +- .../Topology/ContinuousMap/Weierstrass.lean | 2 +- Mathlib/Topology/Homotopy/HomotopyGroup.lean | 19 +++++++++---------- Mathlib/Topology/Sets/Opens.lean | 8 ++++---- 11 files changed, 38 insertions(+), 39 deletions(-) diff --git a/Mathlib/MeasureTheory/Measure/Content.lean b/Mathlib/MeasureTheory/Measure/Content.lean index b07a09b9ad57e..77169173b08ad 100644 --- a/Mathlib/MeasureTheory/Measure/Content.lean +++ b/Mathlib/MeasureTheory/Measure/Content.lean @@ -190,7 +190,7 @@ theorem innerContent_iUnion_nat [R1Space G] ⦃U : ℕ → Set G⦄ rwa [Opens.iSup_def] at this theorem innerContent_comap (f : G ≃ₜ G) (h : ∀ ⦃K : Compacts G⦄, μ (K.map f f.continuous) = μ K) - (U : Opens G) : μ.innerContent (Opens.comap f.toContinuousMap U) = μ.innerContent U := by + (U : Opens G) : μ.innerContent (Opens.comap f U) = μ.innerContent U := by refine (Compacts.equiv f).surjective.iSup_congr _ fun K => iSup_congr_Prop image_subset_iff ?_ intro hK simp only [Equiv.coe_fn_mk, Subtype.mk_eq_mk, Compacts.equiv] @@ -200,7 +200,7 @@ theorem innerContent_comap (f : G ≃ₜ G) (h : ∀ ⦃K : Compacts G⦄, μ (K theorem is_mul_left_invariant_innerContent [Group G] [TopologicalGroup G] (h : ∀ (g : G) {K : Compacts G}, μ (K.map _ <| continuous_mul_left g) = μ K) (g : G) (U : Opens G) : - μ.innerContent (Opens.comap (Homeomorph.mulLeft g).toContinuousMap U) = μ.innerContent U := by + μ.innerContent (Opens.comap (Homeomorph.mulLeft g) U) = μ.innerContent U := by convert μ.innerContent_comap (Homeomorph.mulLeft g) (fun K => h g) U @[to_additive] @@ -211,7 +211,7 @@ theorem innerContent_pos_of_is_mul_left_invariant [Group G] [TopologicalGroup G] rcases compact_covered_by_mul_left_translates K.2 this with ⟨s, hs⟩ suffices μ K ≤ s.card * μ.innerContent U by exact (ENNReal.mul_pos_iff.mp <| hK.bot_lt.trans_le this).2 - have : (K : Set G) ⊆ ↑(⨆ g ∈ s, Opens.comap (Homeomorph.mulLeft g).toContinuousMap U) := by + have : (K : Set G) ⊆ ↑(⨆ g ∈ s, Opens.comap (Homeomorph.mulLeft g : C(G, G)) U) := by simpa only [Opens.iSup_def, Opens.coe_comap, Subtype.coe_mk] refine (μ.le_innerContent _ _ this).trans ?_ refine diff --git a/Mathlib/Topology/Category/TopCat/Basic.lean b/Mathlib/Topology/Category/TopCat/Basic.lean index 8cf13a2f355f4..f1a2371e9013c 100644 --- a/Mathlib/Topology/Category/TopCat/Basic.lean +++ b/Mathlib/Topology/Category/TopCat/Basic.lean @@ -126,8 +126,8 @@ def trivial : Type u ⥤ TopCat.{u} where @[simps] def isoOfHomeo {X Y : TopCat.{u}} (f : X ≃ₜ Y) : X ≅ Y where -- Porting note: previously ⟨f⟩ for hom (inv) and tidy closed proofs - hom := f.toContinuousMap - inv := f.symm.toContinuousMap + hom := (f : C(X, Y)) + inv := (f.symm : C(Y, X)) hom_inv_id := by ext; exact f.symm_apply_apply _ inv_hom_id := by ext; exact f.apply_symm_apply _ diff --git a/Mathlib/Topology/ContinuousMap/Algebra.lean b/Mathlib/Topology/ContinuousMap/Algebra.lean index b6ff6659b4878..59525a680230d 100644 --- a/Mathlib/Topology/ContinuousMap/Algebra.lean +++ b/Mathlib/Topology/ContinuousMap/Algebra.lean @@ -1010,7 +1010,7 @@ variable [ContinuousStar A] [Algebra 𝕜 A] actually a homeomorphism. -/ @[simps] def compStarAlgEquiv' (f : X ≃ₜ Y) : C(Y, A) ≃⋆ₐ[𝕜] C(X, A) := - { f.toContinuousMap.compStarAlgHom' 𝕜 A with + { (f : C(X, Y)).compStarAlgHom' 𝕜 A with toFun := (f : C(X, Y)).compStarAlgHom' 𝕜 A invFun := (f.symm : C(Y, X)).compStarAlgHom' 𝕜 A left_inv := fun g => by @@ -1019,7 +1019,7 @@ def compStarAlgEquiv' (f : X ≃ₜ Y) : C(Y, A) ≃⋆ₐ[𝕜] C(X, A) := right_inv := fun g => by simp only [ContinuousMap.compStarAlgHom'_apply, ContinuousMap.comp_assoc, symm_comp_toContinuousMap, ContinuousMap.comp_id] - map_smul' := fun k a => map_smul (f.toContinuousMap.compStarAlgHom' 𝕜 A) k a } + map_smul' := fun k a => map_smul ((f : C(X, Y)).compStarAlgHom' 𝕜 A) k a } end Homeomorph diff --git a/Mathlib/Topology/ContinuousMap/Basic.lean b/Mathlib/Topology/ContinuousMap/Basic.lean index bd8aca0de948b..69e4f1e0ee255 100644 --- a/Mathlib/Topology/ContinuousMap/Basic.lean +++ b/Mathlib/Topology/ContinuousMap/Basic.lean @@ -429,19 +429,15 @@ namespace Homeomorph variable {α β γ : Type*} [TopologicalSpace α] [TopologicalSpace β] [TopologicalSpace γ] variable (f : α ≃ₜ β) (g : β ≃ₜ γ) +instance instContinuousMapClass : ContinuousMapClass (α ≃ₜ β) α β where + map_continuous f := f.continuous_toFun + /-- The forward direction of a homeomorphism, as a bundled continuous map. -/ -@[simps] -def toContinuousMap (e : α ≃ₜ β) : C(α, β) := +@[simps, deprecated _root_.toContinuousMap (since := "2024-10-12")] +protected def toContinuousMap (e : α ≃ₜ β) : C(α, β) := ⟨e, e.continuous_toFun⟩ -/-- `Homeomorph.toContinuousMap` as a coercion. -/ -instance : Coe (α ≃ₜ β) C(α, β) := - ⟨Homeomorph.toContinuousMap⟩ - --- Porting note: Syntactic tautology -/- theorem toContinuousMap_as_coe : f.toContinuousMap = f := - rfl --/ +attribute [deprecated ContinuousMap.coe_apply (since := "2024-10-12")] toContinuousMap_apply @[simp] theorem coe_refl : (Homeomorph.refl α : C(α, α)) = ContinuousMap.id α := diff --git a/Mathlib/Topology/ContinuousMap/Compact.lean b/Mathlib/Topology/ContinuousMap/Compact.lean index 552c3e592bff1..e5b573631c838 100644 --- a/Mathlib/Topology/ContinuousMap/Compact.lean +++ b/Mathlib/Topology/ContinuousMap/Compact.lean @@ -455,8 +455,8 @@ theorem compRightContinuousMap_apply {X Y : Type*} (T : Type*) [TopologicalSpace def compRightHomeomorph {X Y : Type*} (T : Type*) [TopologicalSpace X] [CompactSpace X] [TopologicalSpace Y] [CompactSpace Y] [PseudoMetricSpace T] (f : X ≃ₜ Y) : C(Y, T) ≃ₜ C(X, T) where - toFun := compRightContinuousMap T f.toContinuousMap - invFun := compRightContinuousMap T f.symm.toContinuousMap + toFun := compRightContinuousMap T f + invFun := compRightContinuousMap T f.symm left_inv g := ext fun _ => congr_arg g (f.apply_symm_apply _) right_inv g := ext fun _ => congr_arg g (f.symm_apply_apply _) diff --git a/Mathlib/Topology/ContinuousMap/ContinuousMapZero.lean b/Mathlib/Topology/ContinuousMap/ContinuousMapZero.lean index 575b4e26ad07c..6d699f5ba3042 100644 --- a/Mathlib/Topology/ContinuousMap/ContinuousMapZero.lean +++ b/Mathlib/Topology/ContinuousMap/ContinuousMapZero.lean @@ -303,15 +303,15 @@ alias uniformEmbedding_comp := isUniformEmbedding_comp sending `0 : X` to `0 : Y`. -/ def _root_.UniformEquiv.arrowCongrLeft₀ {Y : Type*} [TopologicalSpace Y] [Zero Y] (f : X ≃ₜ Y) (hf : f 0 = 0) : C(X, R)₀ ≃ᵤ C(Y, R)₀ where - toFun g := g.comp ⟨f.symm.toContinuousMap, (f.toEquiv.apply_eq_iff_eq_symm_apply.eq ▸ hf).symm⟩ - invFun g := g.comp ⟨f.toContinuousMap, hf⟩ + toFun g := g.comp ⟨f.symm, (f.toEquiv.apply_eq_iff_eq_symm_apply.eq ▸ hf).symm⟩ + invFun g := g.comp ⟨f, hf⟩ left_inv g := ext fun _ ↦ congrArg g <| f.left_inv _ right_inv g := ext fun _ ↦ congrArg g <| f.right_inv _ uniformContinuous_toFun := isUniformEmbedding_toContinuousMap.uniformContinuous_iff.mpr <| - ContinuousMap.uniformContinuous_comp_left f.symm.toContinuousMap |>.comp + ContinuousMap.uniformContinuous_comp_left (f.symm : C(Y, X)) |>.comp isUniformEmbedding_toContinuousMap.uniformContinuous uniformContinuous_invFun := isUniformEmbedding_toContinuousMap.uniformContinuous_iff.mpr <| - ContinuousMap.uniformContinuous_comp_left f.toContinuousMap |>.comp + ContinuousMap.uniformContinuous_comp_left (f : C(X, Y)) |>.comp isUniformEmbedding_toContinuousMap.uniformContinuous end UniformSpace diff --git a/Mathlib/Topology/ContinuousMap/Defs.lean b/Mathlib/Topology/ContinuousMap/Defs.lean index f576214a27758..e1ee866070f68 100644 --- a/Mathlib/Topology/ContinuousMap/Defs.lean +++ b/Mathlib/Topology/ContinuousMap/Defs.lean @@ -93,6 +93,10 @@ protected theorem coe_coe {F : Type*} [FunLike F X Y] [ContinuousMapClass F X Y] ⇑(f : C(X, Y)) = f := rfl +protected theorem coe_apply {F : Type*} [FunLike F X Y] [ContinuousMapClass F X Y] (f : F) (x : X) : + (f : C(X, Y)) x = f x := + rfl + @[ext] theorem ext {f g : C(X, Y)} (h : ∀ a, f a = g a) : f = g := DFunLike.ext _ _ h diff --git a/Mathlib/Topology/ContinuousMap/Polynomial.lean b/Mathlib/Topology/ContinuousMap/Polynomial.lean index 83a3dd5e6a0ed..237d38cd410f8 100644 --- a/Mathlib/Topology/ContinuousMap/Polynomial.lean +++ b/Mathlib/Topology/ContinuousMap/Polynomial.lean @@ -164,7 +164,7 @@ open ContinuousMap /-- The preimage of polynomials on `[0,1]` under the pullback map by `x ↦ (b-a) * x + a` is the polynomials on `[a,b]`. -/ theorem polynomialFunctions.comap_compRightAlgHom_iccHomeoI (a b : ℝ) (h : a < b) : - (polynomialFunctions I).comap (compRightAlgHom ℝ ℝ (iccHomeoI a b h).symm.toContinuousMap) = + (polynomialFunctions I).comap (compRightAlgHom ℝ ℝ (iccHomeoI a b h).symm) = polynomialFunctions (Set.Icc a b) := by ext f fconstructor diff --git a/Mathlib/Topology/ContinuousMap/Weierstrass.lean b/Mathlib/Topology/ContinuousMap/Weierstrass.lean index fbfdd66fb2023..8e903e39101e1 100644 --- a/Mathlib/Topology/ContinuousMap/Weierstrass.lean +++ b/Mathlib/Topology/ContinuousMap/Weierstrass.lean @@ -55,7 +55,7 @@ theorem polynomialFunctions_closure_eq_top (a b : ℝ) : · -- We can pullback continuous functions on `[a,b]` to continuous functions on `[0,1]`, -- by precomposing with an affine map. let W : C(Set.Icc a b, ℝ) →ₐ[ℝ] C(I, ℝ) := - compRightAlgHom ℝ ℝ (iccHomeoI a b h).symm.toContinuousMap + compRightAlgHom ℝ ℝ (iccHomeoI a b h).symm -- This operation is itself a homeomorphism -- (with respect to the norm topologies on continuous functions). let W' : C(Set.Icc a b, ℝ) ≃ₜ C(I, ℝ) := compRightHomeomorph ℝ (iccHomeoI a b h).symm diff --git a/Mathlib/Topology/Homotopy/HomotopyGroup.lean b/Mathlib/Topology/Homotopy/HomotopyGroup.lean index a9ec3d6dc3be4..21ad316e36339 100644 --- a/Mathlib/Topology/Homotopy/HomotopyGroup.lean +++ b/Mathlib/Topology/Homotopy/HomotopyGroup.lean @@ -180,7 +180,7 @@ variable [DecidableEq N] @[simps] def toLoop (i : N) (p : Ω^ N X x) : Ω (Ω^ { j // j ≠ i } X x) const where toFun t := - ⟨(p.val.comp (Cube.insertAt i).toContinuousMap).curry t, fun y yH => + ⟨(p.val.comp (Cube.insertAt i)).curry t, fun y yH => p.property (Cube.insertAt i (t, y)) (Cube.insertAt_boundary i <| Or.inr yH)⟩ source' := by ext t; refine p.property (Cube.insertAt i (0, t)) ⟨i, Or.inl ?_⟩; simp target' := by ext t; refine p.property (Cube.insertAt i (1, t)) ⟨i, Or.inr ?_⟩; simp @@ -200,10 +200,10 @@ theorem continuous_toLoop (i : N) : Continuous (@toLoop N X _ x _ i) := @[simps] def fromLoop (i : N) (p : Ω (Ω^ { j // j ≠ i } X x) const) : Ω^ N X x := ⟨(ContinuousMap.comp ⟨Subtype.val, by fun_prop⟩ p.toContinuousMap).uncurry.comp - (Cube.splitAt i).toContinuousMap, + (Cube.splitAt i), by rintro y ⟨j, Hj⟩ - simp only [ContinuousMap.comp_apply, toContinuousMap_apply, + simp only [ContinuousMap.comp_apply, ContinuousMap.coe_coe, funSplitAt_apply, ContinuousMap.uncurry_apply, ContinuousMap.coe_mk, Function.uncurry_apply_pair] obtain rfl | Hne := eq_or_ne j i @@ -243,8 +243,8 @@ theorem fromLoop_apply (i : N) {p : Ω (Ω^ { j // j ≠ i } X x) const} {t : I^ /-- Composition with `Cube.insertAt` as a continuous map. -/ abbrev cCompInsert (i : N) : C(C(I^N, X), C(I × I^{ j // j ≠ i }, X)) := - ⟨fun f => f.comp (Cube.insertAt i).toContinuousMap, - (Cube.insertAt i).toContinuousMap.continuous_comp_left⟩ + ⟨fun f => f.comp (Cube.insertAt i), + (toContinuousMap <| Cube.insertAt i).continuous_comp_left⟩ /-- A homotopy between `n+1`-dimensional loops `p` and `q` constant on the boundary seen as a homotopy between two paths in the space of `n`-dimensional paths. -/ @@ -280,7 +280,7 @@ theorem homotopicTo (i : N) {p q : Ω^ N X x} : C(I × I^N, X) := (ContinuousMap.comp ⟨_, ContinuousMap.continuous_uncurry⟩ (ContinuousMap.comp ⟨Subtype.val, by continuity⟩ H.toContinuousMap).curry).uncurry.comp <| - (ContinuousMap.id I).prodMap (Cube.splitAt i).toContinuousMap + (ContinuousMap.id I).prodMap (Cube.splitAt i) theorem homotopicFrom (i : N) {p q : Ω^ N X x} : (toLoop i p).Homotopic (toLoop i q) → Homotopic p q := by @@ -289,15 +289,14 @@ theorem homotopicFrom (i : N) {p q : Ω^ N X x} : · rintro t y ⟨j, jH⟩ erw [homotopyFrom_apply] obtain rfl | h := eq_or_ne j i - · simp only [Prod.map_apply, id_eq, toContinuousMap_apply, funSplitAt_apply, - Function.uncurry_apply_pair] + · simp only [Prod.map_apply, id_eq, funSplitAt_apply, Function.uncurry_apply_pair] rw [H.eq_fst] exacts [congr_arg p ((Cube.splitAt j).left_inv _), jH] · rw [p.2 _ ⟨j, jH⟩]; apply boundary; exact ⟨⟨j, h⟩, jH⟩ all_goals intro apply (homotopyFrom_apply _ _ _).trans - simp only [Prod.map_apply, id_eq, toContinuousMap_apply, funSplitAt_apply, + simp only [Prod.map_apply, id_eq, funSplitAt_apply, Function.uncurry_apply_pair, ContinuousMap.HomotopyWith.apply_zero, ContinuousMap.HomotopyWith.apply_one, ne_eq, Path.coe_toContinuousMap, toLoop_apply_coe, ContinuousMap.curry_apply, ContinuousMap.comp_apply] @@ -315,7 +314,7 @@ def transAt (i : N) (f g : Ω^ N X x) : Ω^ N X x := (by ext1; symm dsimp only [Path.trans, fromLoop, Path.coe_mk_mk, Function.comp_apply, mk_apply, - ContinuousMap.comp_apply, toContinuousMap_apply, funSplitAt_apply, + ContinuousMap.comp_apply, ContinuousMap.coe_coe, funSplitAt_apply, ContinuousMap.uncurry_apply, ContinuousMap.coe_mk, Function.uncurry_apply_pair] split_ifs · show f _ = _; congr 1 diff --git a/Mathlib/Topology/Sets/Opens.lean b/Mathlib/Topology/Sets/Opens.lean index 1b2ffd4204a29..bd142710c0990 100644 --- a/Mathlib/Topology/Sets/Opens.lean +++ b/Mathlib/Topology/Sets/Opens.lean @@ -379,10 +379,10 @@ theorem comap_injective [T0Space β] : Injective (comap : C(α, β) → FrameHom /-- A homeomorphism induces an order-preserving equivalence on open sets, by taking comaps. -/ @[simps (config := .asFn) apply] def _root_.Homeomorph.opensCongr (f : α ≃ₜ β) : Opens α ≃o Opens β where - toFun := Opens.comap f.symm.toContinuousMap - invFun := Opens.comap f.toContinuousMap - left_inv := fun _ => ext <| f.toEquiv.preimage_symm_preimage _ - right_inv := fun _ => ext <| f.toEquiv.symm_preimage_preimage _ + toFun := Opens.comap (f.symm : C(β, α)) + invFun := Opens.comap (f : C(α, β)) + left_inv _ := ext <| f.toEquiv.preimage_symm_preimage _ + right_inv _ := ext <| f.toEquiv.symm_preimage_preimage _ map_rel_iff' := by simp only [← SetLike.coe_subset_coe]; exact f.symm.surjective.preimage_subset_preimage_iff From a6bc4c5ee86f2ccb63f6d63406cd3db02d3de29a Mon Sep 17 00:00:00 2001 From: sgouezel Date: Fri, 18 Oct 2024 11:08:02 +0000 Subject: [PATCH 139/505] chore: reorganize the `ContinuousOn` file (#17901) Before adding some API for `ContinuousWithinAt` to have it match the one for `DifferentiableWithinAt`, and `ContDiffWithinAt`, and manifold versions, I'd rather start from a clean file. The file was a mess of random things added at random places, I tried to make it coherent. Nothing has been added or removed, no name has been changed, pure shuffling around. --- Mathlib/Topology/Constructions.lean | 7 + Mathlib/Topology/ContinuousOn.lean | 732 +++++++++++++++------------- 2 files changed, 398 insertions(+), 341 deletions(-) diff --git a/Mathlib/Topology/Constructions.lean b/Mathlib/Topology/Constructions.lean index e0bb583947c37..82bce23e4c0c9 100644 --- a/Mathlib/Topology/Constructions.lean +++ b/Mathlib/Topology/Constructions.lean @@ -1134,6 +1134,13 @@ lemma isClosed_preimage_val {s t : Set X} : IsClosed (s ↓∩ t) ↔ s ∩ clos Subtype.image_preimage_coe, subset_antisymm_iff, and_iff_left, Set.subset_inter_iff, and_iff_right] exacts [Set.inter_subset_left, Set.subset_inter Set.inter_subset_left subset_closure] + +theorem frontier_inter_open_inter {s t : Set X} (ht : IsOpen t) : + frontier (s ∩ t) ∩ t = frontier s ∩ t := by + simp only [Set.inter_comm _ t, ← Subtype.preimage_coe_eq_preimage_coe_iff, + ht.isOpenMap_subtype_val.preimage_frontier_eq_frontier_preimage continuous_subtype_val, + Subtype.preimage_coe_self_inter] + end Subtype section Quotient diff --git a/Mathlib/Topology/ContinuousOn.lean b/Mathlib/Topology/ContinuousOn.lean index a09c8362c391f..d69f7aec79123 100644 --- a/Mathlib/Topology/ContinuousOn.lean +++ b/Mathlib/Topology/ContinuousOn.lean @@ -9,13 +9,14 @@ import Mathlib.Topology.Constructions /-! # Neighborhoods and continuity relative to a subset -This file defines relative versions +This file develops API on the relative versions * `nhdsWithin` of `nhds` * `ContinuousOn` of `Continuous` * `ContinuousWithinAt` of `ContinuousAt` -and proves their basic properties, including the relationships between +related to continuity, which are defined in previous definition files. +Their basic properties studied in this file include the relationships between these restricted notions and the corresponding notions for the subtype equipped with the subspace topology. @@ -32,6 +33,10 @@ open Set Filter Function Topology Filter variable {α : Type*} {β : Type*} {γ : Type*} {δ : Type*} variable [TopologicalSpace α] +/-! +## Properties of the neighborhood-within filter +-/ + @[simp] theorem nhds_bind_nhdsWithin {a : α} {s : Set α} : ((𝓝 a).bind fun x => 𝓝[s] x) = 𝓝[s] a := bind_inf_principal.trans <| congr_arg₂ _ nhds_bind_nhds rfl @@ -299,6 +304,8 @@ instance Pi.instNeBotNhdsWithinIoi [Nonempty ι] [∀ i, Preorder (π i)] {x : [∀ i, (𝓝[>] x i).NeBot] : (𝓝[>] x).NeBot := Pi.instNeBotNhdsWithinIio (π := fun i ↦ (π i)ᵒᵈ) (x := fun i ↦ OrderDual.toDual (x i)) +end Pi + theorem Filter.Tendsto.piecewise_nhdsWithin {f g : α → β} {t : Set α} [∀ x, Decidable (x ∈ t)] {a : α} {s : Set α} {l : Filter β} (h₀ : Tendsto f (𝓝[s ∩ t] a) l) (h₁ : Tendsto g (𝓝[s ∩ tᶜ] a) l) : Tendsto (piecewise t f g) (𝓝[s] a) l := by @@ -438,8 +445,16 @@ theorem tendsto_nhdsWithin_iff_subtype {s : Set α} {a : α} (h : a ∈ s) (f : Tendsto f (𝓝[s] a) l ↔ Tendsto (s.restrict f) (𝓝 ⟨a, h⟩) l := by rw [nhdsWithin_eq_map_subtype_coe h, tendsto_map'_iff]; rfl +/-! +## Local continuity properties of functions +-/ + variable [TopologicalSpace β] [TopologicalSpace γ] [TopologicalSpace δ] +/-! +### `ContinuousWithinAt` +-/ + /-- If a function is continuous within `s` at `x`, then it tends to `f x` within `s` by definition. We register this fact for use with the dot notation, especially to use `Filter.Tendsto.comp` as `ContinuousWithinAt.comp` will have a different meaning. -/ @@ -447,10 +462,6 @@ theorem ContinuousWithinAt.tendsto {f : α → β} {s : Set α} {x : α} (h : Co Tendsto f (𝓝[s] x) (𝓝 (f x)) := h -theorem ContinuousOn.continuousWithinAt {f : α → β} {s : Set α} {x : α} (hf : ContinuousOn f s) - (hx : x ∈ s) : ContinuousWithinAt f s x := - hf x hx - theorem continuousWithinAt_univ (f : α → β) (x : α) : ContinuousWithinAt f Set.univ x ↔ ContinuousAt f x := by rw [ContinuousAt, ContinuousWithinAt, nhdsWithin_univ] @@ -471,93 +482,44 @@ theorem ContinuousWithinAt.tendsto_nhdsWithin_image {f : α → β} {x : α} {s (h : ContinuousWithinAt f s x) : Tendsto f (𝓝[s] x) (𝓝[f '' s] f x) := h.tendsto_nhdsWithin (mapsTo_image _ _) -theorem ContinuousWithinAt.prod_map {f : α → γ} {g : β → δ} {s : Set α} {t : Set β} {x : α} {y : β} - (hf : ContinuousWithinAt f s x) (hg : ContinuousWithinAt g t y) : - ContinuousWithinAt (Prod.map f g) (s ×ˢ t) (x, y) := by - unfold ContinuousWithinAt at * - rw [nhdsWithin_prod_eq, Prod.map, nhds_prod_eq] - exact hf.prod_map hg - -theorem continuousWithinAt_prod_of_discrete_left [DiscreteTopology α] - {f : α × β → γ} {s : Set (α × β)} {x : α × β} : - ContinuousWithinAt f s x ↔ ContinuousWithinAt (f ⟨x.1, ·⟩) {b | (x.1, b) ∈ s} x.2 := by - rw [← x.eta]; simp_rw [ContinuousWithinAt, nhdsWithin, nhds_prod_eq, nhds_discrete, pure_prod, - ← map_inf_principal_preimage]; rfl - -theorem continuousWithinAt_prod_of_discrete_right [DiscreteTopology β] - {f : α × β → γ} {s : Set (α × β)} {x : α × β} : - ContinuousWithinAt f s x ↔ ContinuousWithinAt (f ⟨·, x.2⟩) {a | (a, x.2) ∈ s} x.1 := by - rw [← x.eta]; simp_rw [ContinuousWithinAt, nhdsWithin, nhds_prod_eq, nhds_discrete, prod_pure, - ← map_inf_principal_preimage]; rfl - -theorem continuousAt_prod_of_discrete_left [DiscreteTopology α] {f : α × β → γ} {x : α × β} : - ContinuousAt f x ↔ ContinuousAt (f ⟨x.1, ·⟩) x.2 := by - simp_rw [← continuousWithinAt_univ]; exact continuousWithinAt_prod_of_discrete_left - -theorem continuousAt_prod_of_discrete_right [DiscreteTopology β] {f : α × β → γ} {x : α × β} : - ContinuousAt f x ↔ ContinuousAt (f ⟨·, x.2⟩) x.1 := by - simp_rw [← continuousWithinAt_univ]; exact continuousWithinAt_prod_of_discrete_right - -theorem continuousOn_prod_of_discrete_left [DiscreteTopology α] {f : α × β → γ} {s : Set (α × β)} : - ContinuousOn f s ↔ ∀ a, ContinuousOn (f ⟨a, ·⟩) {b | (a, b) ∈ s} := by - simp_rw [ContinuousOn, Prod.forall, continuousWithinAt_prod_of_discrete_left]; rfl - -theorem continuousOn_prod_of_discrete_right [DiscreteTopology β] {f : α × β → γ} {s : Set (α × β)} : - ContinuousOn f s ↔ ∀ b, ContinuousOn (f ⟨·, b⟩) {a | (a, b) ∈ s} := by - simp_rw [ContinuousOn, Prod.forall, continuousWithinAt_prod_of_discrete_right]; apply forall_swap - -/-- If a function `f a b` is such that `y ↦ f a b` is continuous for all `a`, and `a` lives in a -discrete space, then `f` is continuous, and vice versa. -/ -theorem continuous_prod_of_discrete_left [DiscreteTopology α] {f : α × β → γ} : - Continuous f ↔ ∀ a, Continuous (f ⟨a, ·⟩) := by - simp_rw [continuous_iff_continuousOn_univ]; exact continuousOn_prod_of_discrete_left - -theorem continuous_prod_of_discrete_right [DiscreteTopology β] {f : α × β → γ} : - Continuous f ↔ ∀ b, Continuous (f ⟨·, b⟩) := by - simp_rw [continuous_iff_continuousOn_univ]; exact continuousOn_prod_of_discrete_right - -theorem isOpenMap_prod_of_discrete_left [DiscreteTopology α] {f : α × β → γ} : - IsOpenMap f ↔ ∀ a, IsOpenMap (f ⟨a, ·⟩) := by - simp_rw [isOpenMap_iff_nhds_le, Prod.forall, nhds_prod_eq, nhds_discrete, pure_prod, map_map] - rfl - -theorem isOpenMap_prod_of_discrete_right [DiscreteTopology β] {f : α × β → γ} : - IsOpenMap f ↔ ∀ b, IsOpenMap (f ⟨·, b⟩) := by - simp_rw [isOpenMap_iff_nhds_le, Prod.forall, forall_swap (α := α) (β := β), nhds_prod_eq, - nhds_discrete, prod_pure, map_map]; rfl +theorem nhdsWithin_le_comap {x : α} {s : Set α} {f : α → β} (ctsf : ContinuousWithinAt f s x) : + 𝓝[s] x ≤ comap f (𝓝[f '' s] f x) := + ctsf.tendsto_nhdsWithin_image.le_comap -theorem continuousWithinAt_pi {ι : Type*} {π : ι → Type*} [∀ i, TopologicalSpace (π i)] - {f : α → ∀ i, π i} {s : Set α} {x : α} : - ContinuousWithinAt f s x ↔ ∀ i, ContinuousWithinAt (fun y => f y i) s x := - tendsto_pi_nhds +theorem ContinuousWithinAt.preimage_mem_nhdsWithin {f : α → β} {x : α} {s : Set α} {t : Set β} + (h : ContinuousWithinAt f s x) (ht : t ∈ 𝓝 (f x)) : f ⁻¹' t ∈ 𝓝[s] x := + h ht -theorem continuousOn_pi {ι : Type*} {π : ι → Type*} [∀ i, TopologicalSpace (π i)] - {f : α → ∀ i, π i} {s : Set α} : ContinuousOn f s ↔ ∀ i, ContinuousOn (fun y => f y i) s := - ⟨fun h i x hx => tendsto_pi_nhds.1 (h x hx) i, fun h x hx => tendsto_pi_nhds.2 fun i => h i x hx⟩ +theorem ContinuousWithinAt.preimage_mem_nhdsWithin' {f : α → β} {x : α} {s : Set α} {t : Set β} + (h : ContinuousWithinAt f s x) (ht : t ∈ 𝓝[f '' s] f x) : f ⁻¹' t ∈ 𝓝[s] x := + h.tendsto_nhdsWithin (mapsTo_image _ _) ht -@[fun_prop] -theorem continuousOn_pi' {ι : Type*} {π : ι → Type*} [∀ i, TopologicalSpace (π i)] - {f : α → ∀ i, π i} {s : Set α} (hf : ∀ i, ContinuousOn (fun y => f y i) s) : - ContinuousOn f s := - continuousOn_pi.2 hf +theorem ContinuousWithinAt.preimage_mem_nhdsWithin'' + {f : α → β} {x : α} {y : β} {s t : Set β} + (h : ContinuousWithinAt f (f ⁻¹' s) x) (ht : t ∈ 𝓝[s] y) (hxy : y = f x) : + f ⁻¹' t ∈ 𝓝[f ⁻¹' s] x := by + rw [hxy] at ht + exact h.preimage_mem_nhdsWithin' (nhdsWithin_mono _ (image_preimage_subset f s) ht) -theorem ContinuousWithinAt.fin_insertNth {n} {π : Fin (n + 1) → Type*} - [∀ i, TopologicalSpace (π i)] (i : Fin (n + 1)) {f : α → π i} {a : α} {s : Set α} - (hf : ContinuousWithinAt f s a) {g : α → ∀ j : Fin n, π (i.succAbove j)} - (hg : ContinuousWithinAt g s a) : ContinuousWithinAt (fun a => i.insertNth (f a) (g a)) s a := - hf.tendsto.fin_insertNth i hg +theorem continuousWithinAt_of_not_mem_closure {f : α → β} {s : Set α} {x : α} (hx : x ∉ closure s) : + ContinuousWithinAt f s x := by + rw [mem_closure_iff_nhdsWithin_neBot, not_neBot] at hx + rw [ContinuousWithinAt, hx] + exact tendsto_bot -nonrec theorem ContinuousOn.fin_insertNth {n} {π : Fin (n + 1) → Type*} - [∀ i, TopologicalSpace (π i)] (i : Fin (n + 1)) {f : α → π i} {s : Set α} - (hf : ContinuousOn f s) {g : α → ∀ j : Fin n, π (i.succAbove j)} (hg : ContinuousOn g s) : - ContinuousOn (fun a => i.insertNth (f a) (g a)) s := fun a ha => - (hf a ha).fin_insertNth i (hg a ha) +/-! +### `ContinuousOn` +-/ theorem continuousOn_iff {f : α → β} {s : Set α} : ContinuousOn f s ↔ ∀ x ∈ s, ∀ t : Set β, IsOpen t → f x ∈ t → ∃ u, IsOpen u ∧ x ∈ u ∧ u ∩ s ⊆ f ⁻¹' t := by simp only [ContinuousOn, ContinuousWithinAt, tendsto_nhds, mem_nhdsWithin] +theorem ContinuousOn.continuousWithinAt {f : α → β} {s : Set α} {x : α} (hf : ContinuousOn f s) + (hx : x ∈ s) : ContinuousWithinAt f s x := + hf x hx + theorem continuousOn_iff_continuous_restrict {f : α → β} {s : Set α} : ContinuousOn f s ↔ Continuous (s.restrict f) := by rw [ContinuousOn, continuous_iff_continuousAt]; constructor @@ -606,10 +568,6 @@ theorem continuousOn_iff_isClosed {f : α → β} {s : Set α} : simp only [Subtype.preimage_coe_eq_preimage_coe_iff, eq_comm, Set.inter_comm s] rw [continuousOn_iff_continuous_restrict, continuous_iff_isClosed]; simp only [this] -theorem ContinuousOn.prod_map {f : α → γ} {g : β → δ} {s : Set α} {t : Set β} - (hf : ContinuousOn f s) (hg : ContinuousOn g t) : ContinuousOn (Prod.map f g) (s ×ˢ t) := - fun ⟨x, y⟩ ⟨hx, hy⟩ => ContinuousWithinAt.prod_map (hf x hx) (hg y hy) - theorem continuous_of_cover_nhds {ι : Sort*} {f : α → β} {s : ι → Set α} (hs : ∀ x : α, ∃ i, s i ∈ 𝓝 x) (hf : ∀ i, ContinuousOn f (s i)) : Continuous f := @@ -629,9 +587,66 @@ theorem Set.Subsingleton.continuousOn {s : Set α} (hs : s.Subsingleton) (f : α ContinuousOn f s := hs.induction_on (continuousOn_empty f) (continuousOn_singleton f) -theorem nhdsWithin_le_comap {x : α} {s : Set α} {f : α → β} (ctsf : ContinuousWithinAt f s x) : - 𝓝[s] x ≤ comap f (𝓝[f '' s] f x) := - ctsf.tendsto_nhdsWithin_image.le_comap +theorem continuousOn_open_iff {f : α → β} {s : Set α} (hs : IsOpen s) : + ContinuousOn f s ↔ ∀ t, IsOpen t → IsOpen (s ∩ f ⁻¹' t) := by + rw [continuousOn_iff'] + constructor + · intro h t ht + rcases h t ht with ⟨u, u_open, hu⟩ + rw [inter_comm, hu] + apply IsOpen.inter u_open hs + · intro h t ht + refine ⟨s ∩ f ⁻¹' t, h t ht, ?_⟩ + rw [@inter_comm _ s (f ⁻¹' t), inter_assoc, inter_self] + +theorem ContinuousOn.isOpen_inter_preimage {f : α → β} {s : Set α} {t : Set β} + (hf : ContinuousOn f s) (hs : IsOpen s) (ht : IsOpen t) : IsOpen (s ∩ f ⁻¹' t) := + (continuousOn_open_iff hs).1 hf t ht + +theorem ContinuousOn.isOpen_preimage {f : α → β} {s : Set α} {t : Set β} (h : ContinuousOn f s) + (hs : IsOpen s) (hp : f ⁻¹' t ⊆ s) (ht : IsOpen t) : IsOpen (f ⁻¹' t) := by + convert (continuousOn_open_iff hs).mp h t ht + rw [inter_comm, inter_eq_self_of_subset_left hp] + +theorem ContinuousOn.preimage_isClosed_of_isClosed {f : α → β} {s : Set α} {t : Set β} + (hf : ContinuousOn f s) (hs : IsClosed s) (ht : IsClosed t) : IsClosed (s ∩ f ⁻¹' t) := by + rcases continuousOn_iff_isClosed.1 hf t ht with ⟨u, hu⟩ + rw [inter_comm, hu.2] + apply IsClosed.inter hu.1 hs + +theorem ContinuousOn.preimage_interior_subset_interior_preimage {f : α → β} {s : Set α} {t : Set β} + (hf : ContinuousOn f s) (hs : IsOpen s) : s ∩ f ⁻¹' interior t ⊆ s ∩ interior (f ⁻¹' t) := + calc + s ∩ f ⁻¹' interior t ⊆ interior (s ∩ f ⁻¹' t) := + interior_maximal (inter_subset_inter (Subset.refl _) (preimage_mono interior_subset)) + (hf.isOpen_inter_preimage hs isOpen_interior) + _ = s ∩ interior (f ⁻¹' t) := by rw [interior_inter, hs.interior_eq] + +theorem continuousOn_of_locally_continuousOn {f : α → β} {s : Set α} + (h : ∀ x ∈ s, ∃ t, IsOpen t ∧ x ∈ t ∧ ContinuousOn f (s ∩ t)) : ContinuousOn f s := by + intro x xs + rcases h x xs with ⟨t, open_t, xt, ct⟩ + have := ct x ⟨xs, xt⟩ + rwa [ContinuousWithinAt, ← nhdsWithin_restrict _ xt open_t] at this + +theorem continuousOn_to_generateFrom_iff {β} {s : Set α} {T : Set (Set β)} {f : α → β} : + @ContinuousOn α β _ (.generateFrom T) f s ↔ ∀ x ∈ s, ∀ t ∈ T, f x ∈ t → f ⁻¹' t ∈ 𝓝[s] x := + forall₂_congr fun x _ => by + delta ContinuousWithinAt + simp only [TopologicalSpace.nhds_generateFrom, tendsto_iInf, tendsto_principal, mem_setOf_eq, + and_imp] + exact forall_congr' fun t => forall_swap + +-- Porting note: dropped an unneeded assumption +theorem continuousOn_isOpen_of_generateFrom {β : Type*} {s : Set α} {T : Set (Set β)} {f : α → β} + (h : ∀ t ∈ T, IsOpen (s ∩ f ⁻¹' t)) : + @ContinuousOn α β _ (.generateFrom T) f s := + continuousOn_to_generateFrom_iff.2 fun _x hx t ht hxt => mem_nhdsWithin.2 + ⟨_, h t ht, ⟨hx, hxt⟩, fun _y hy => hy.1.2⟩ + +/-! +### Congruence and monotonicity properties with respect to sets +-/ theorem ContinuousWithinAt.mono {f : α → β} {s t : Set α} {x : α} (h : ContinuousWithinAt f t x) (hs : s ⊆ t) : ContinuousWithinAt f s x := @@ -661,31 +676,6 @@ theorem ContinuousWithinAt.union {f : α → β} {s t : Set α} {x : α} (hs : C (ht : ContinuousWithinAt f t x) : ContinuousWithinAt f (s ∪ t) x := continuousWithinAt_union.2 ⟨hs, ht⟩ -theorem ContinuousWithinAt.mem_closure_image {f : α → β} {s : Set α} {x : α} - (h : ContinuousWithinAt f s x) (hx : x ∈ closure s) : f x ∈ closure (f '' s) := - haveI := mem_closure_iff_nhdsWithin_neBot.1 hx - mem_closure_of_tendsto h <| mem_of_superset self_mem_nhdsWithin (subset_preimage_image f s) - -theorem ContinuousWithinAt.mem_closure {f : α → β} {s : Set α} {x : α} {A : Set β} - (h : ContinuousWithinAt f s x) (hx : x ∈ closure s) (hA : MapsTo f s A) : f x ∈ closure A := - closure_mono (image_subset_iff.2 hA) (h.mem_closure_image hx) - -theorem Set.MapsTo.closure_of_continuousWithinAt {f : α → β} {s : Set α} {t : Set β} - (h : MapsTo f s t) (hc : ∀ x ∈ closure s, ContinuousWithinAt f s x) : - MapsTo f (closure s) (closure t) := fun x hx => (hc x hx).mem_closure hx h - -theorem Set.MapsTo.closure_of_continuousOn {f : α → β} {s : Set α} {t : Set β} (h : MapsTo f s t) - (hc : ContinuousOn f (closure s)) : MapsTo f (closure s) (closure t) := - h.closure_of_continuousWithinAt fun x hx => (hc x hx).mono subset_closure - -theorem ContinuousWithinAt.image_closure {f : α → β} {s : Set α} - (hf : ∀ x ∈ closure s, ContinuousWithinAt f s x) : f '' closure s ⊆ closure (f '' s) := - ((mapsTo_image f s).closure_of_continuousWithinAt hf).image_subset - -theorem ContinuousOn.image_closure {f : α → β} {s : Set α} (hf : ContinuousOn f (closure s)) : - f '' closure s ⊆ closure (f '' s) := - ContinuousWithinAt.image_closure fun x hx => (hf x hx).mono subset_closure - @[simp] theorem continuousWithinAt_singleton {f : α → β} {x : α} : ContinuousWithinAt f {x} x := by simp only [ContinuousWithinAt, nhdsWithin_singleton, tendsto_pure_nhds] @@ -712,50 +702,16 @@ theorem continuousWithinAt_compl_self {f : α → β} {a : α} : ContinuousWithinAt f {a}ᶜ a ↔ ContinuousAt f a := by rw [compl_eq_univ_diff, continuousWithinAt_diff_self, continuousWithinAt_univ] -@[simp] -theorem continuousWithinAt_update_same [DecidableEq α] {f : α → β} {s : Set α} {x : α} {y : β} : - ContinuousWithinAt (update f x y) s x ↔ Tendsto f (𝓝[s \ {x}] x) (𝓝 y) := - calc - ContinuousWithinAt (update f x y) s x ↔ Tendsto (update f x y) (𝓝[s \ {x}] x) (𝓝 y) := by - { rw [← continuousWithinAt_diff_self, ContinuousWithinAt, update_same] } - _ ↔ Tendsto f (𝓝[s \ {x}] x) (𝓝 y) := - tendsto_congr' <| eventually_nhdsWithin_iff.2 <| Eventually.of_forall - fun _ hz => update_noteq hz.2 _ _ - -@[simp] -theorem continuousAt_update_same [DecidableEq α] {f : α → β} {x : α} {y : β} : - ContinuousAt (Function.update f x y) x ↔ Tendsto f (𝓝[≠] x) (𝓝 y) := by - rw [← continuousWithinAt_univ, continuousWithinAt_update_same, compl_eq_univ_diff] - -theorem IsOpenMap.continuousOn_image_of_leftInvOn {f : α → β} {s : Set α} - (h : IsOpenMap (s.restrict f)) {finv : β → α} (hleft : LeftInvOn finv f s) : - ContinuousOn finv (f '' s) := by - refine continuousOn_iff'.2 fun t ht => ⟨f '' (t ∩ s), ?_, ?_⟩ - · rw [← image_restrict] - exact h _ (ht.preimage continuous_subtype_val) - · rw [inter_eq_self_of_subset_left (image_subset f inter_subset_right), hleft.image_inter'] - -theorem IsOpenMap.continuousOn_range_of_leftInverse {f : α → β} (hf : IsOpenMap f) {finv : β → α} - (hleft : Function.LeftInverse finv f) : ContinuousOn finv (range f) := by - rw [← image_univ] - exact (hf.restrict isOpen_univ).continuousOn_image_of_leftInvOn fun x _ => hleft x +theorem ContinuousOn.mono {f : α → β} {s t : Set α} (hf : ContinuousOn f s) (h : t ⊆ s) : + ContinuousOn f t := fun x hx => (hf x (h hx)).mono_left (nhdsWithin_mono _ h) -theorem ContinuousOn.congr_mono {f g : α → β} {s s₁ : Set α} (h : ContinuousOn f s) - (h' : EqOn g f s₁) (h₁ : s₁ ⊆ s) : ContinuousOn g s₁ := by - intro x hx - unfold ContinuousWithinAt - have A := (h x (h₁ hx)).mono h₁ - unfold ContinuousWithinAt at A - rw [← h' hx] at A - exact A.congr' h'.eventuallyEq_nhdsWithin.symm +theorem antitone_continuousOn {f : α → β} : Antitone (ContinuousOn f) := fun _s _t hst hf => + hf.mono hst -theorem ContinuousOn.congr {f g : α → β} {s : Set α} (h : ContinuousOn f s) (h' : EqOn g f s) : - ContinuousOn g s := - h.congr_mono h' (Subset.refl _) -theorem continuousOn_congr {f g : α → β} {s : Set α} (h' : EqOn g f s) : - ContinuousOn g s ↔ ContinuousOn f s := - ⟨fun h => ContinuousOn.congr h h'.symm, fun h => h.congr h'⟩ +/-! +### Relation between `ContinuousAt` and `ContinuousWithinAt` +-/ theorem ContinuousAt.continuousWithinAt {f : α → β} {s : Set α} {x : α} (h : ContinuousAt f x) : ContinuousWithinAt f s x := @@ -780,7 +736,61 @@ theorem ContinuousOn.continuousAt {f : α → β} {s : Set α} {x : α} (h : Con theorem ContinuousAt.continuousOn {f : α → β} {s : Set α} (hcont : ∀ x ∈ s, ContinuousAt f x) : ContinuousOn f s := fun x hx => (hcont x hx).continuousWithinAt -theorem ContinuousWithinAt.comp {g : β → γ} {f : α → β} {s : Set α} {t : Set β} {x : α} +@[fun_prop] +theorem Continuous.continuousOn {f : α → β} {s : Set α} (h : Continuous f) : ContinuousOn f s := by + rw [continuous_iff_continuousOn_univ] at h + exact h.mono (subset_univ _) + +theorem Continuous.continuousWithinAt {f : α → β} {s : Set α} {x : α} (h : Continuous f) : + ContinuousWithinAt f s x := + h.continuousAt.continuousWithinAt + + +/-! +### Congruence properties with respect to functions +-/ + +theorem ContinuousOn.congr_mono {f g : α → β} {s s₁ : Set α} (h : ContinuousOn f s) + (h' : EqOn g f s₁) (h₁ : s₁ ⊆ s) : ContinuousOn g s₁ := by + intro x hx + unfold ContinuousWithinAt + have A := (h x (h₁ hx)).mono h₁ + unfold ContinuousWithinAt at A + rw [← h' hx] at A + exact A.congr' h'.eventuallyEq_nhdsWithin.symm + +theorem ContinuousOn.congr {f g : α → β} {s : Set α} (h : ContinuousOn f s) (h' : EqOn g f s) : + ContinuousOn g s := + h.congr_mono h' (Subset.refl _) + +theorem continuousOn_congr {f g : α → β} {s : Set α} (h' : EqOn g f s) : + ContinuousOn g s ↔ ContinuousOn f s := + ⟨fun h => ContinuousOn.congr h h'.symm, fun h => h.congr h'⟩ + +theorem Filter.EventuallyEq.congr_continuousWithinAt {f g : α → β} {s : Set α} {x : α} + (h : f =ᶠ[𝓝[s] x] g) (hx : f x = g x) : + ContinuousWithinAt f s x ↔ ContinuousWithinAt g s x := by + rw [ContinuousWithinAt, hx, tendsto_congr' h, ContinuousWithinAt] + +theorem ContinuousWithinAt.congr_of_eventuallyEq {f f₁ : α → β} {s : Set α} {x : α} + (h : ContinuousWithinAt f s x) (h₁ : f₁ =ᶠ[𝓝[s] x] f) (hx : f₁ x = f x) : + ContinuousWithinAt f₁ s x := + (h₁.congr_continuousWithinAt hx).2 h + +theorem ContinuousWithinAt.congr {f f₁ : α → β} {s : Set α} {x : α} (h : ContinuousWithinAt f s x) + (h₁ : ∀ y ∈ s, f₁ y = f y) (hx : f₁ x = f x) : ContinuousWithinAt f₁ s x := + h.congr_of_eventuallyEq (mem_of_superset self_mem_nhdsWithin h₁) hx + +theorem ContinuousWithinAt.congr_mono {f g : α → β} {s s₁ : Set α} {x : α} + (h : ContinuousWithinAt f s x) (h' : EqOn g f s₁) (h₁ : s₁ ⊆ s) (hx : g x = f x) : + ContinuousWithinAt g s₁ x := + (h.mono h₁).congr h' hx + +/-! +### Composition +-/ + +theorem ContinuousWithinAt.comp {g : β → γ} {f : α → β} {s : Set α} {t : Set β} {x : α} (hg : ContinuousWithinAt g t (f x)) (hf : ContinuousWithinAt f s x) (h : MapsTo f s t) : ContinuousWithinAt (g ∘ f) s x := hg.tendsto.comp (hf.tendsto_nhdsWithin h) @@ -803,26 +813,11 @@ theorem ContinuousOn.comp'' {g : β → γ} {f : α → β} {s : Set α} {t : Se (hf : ContinuousOn f s) (h : Set.MapsTo f s t) : ContinuousOn (fun x => g (f x)) s := ContinuousOn.comp hg hf h -theorem ContinuousOn.mono {f : α → β} {s t : Set α} (hf : ContinuousOn f s) (h : t ⊆ s) : - ContinuousOn f t := fun x hx => (hf x (h hx)).mono_left (nhdsWithin_mono _ h) - -theorem antitone_continuousOn {f : α → β} : Antitone (ContinuousOn f) := fun _s _t hst hf => - hf.mono hst - @[fun_prop] theorem ContinuousOn.comp' {g : β → γ} {f : α → β} {s : Set α} {t : Set β} (hg : ContinuousOn g t) (hf : ContinuousOn f s) : ContinuousOn (g ∘ f) (s ∩ f ⁻¹' t) := hg.comp (hf.mono inter_subset_left) inter_subset_right -@[fun_prop] -theorem Continuous.continuousOn {f : α → β} {s : Set α} (h : Continuous f) : ContinuousOn f s := by - rw [continuous_iff_continuousOn_univ] at h - exact h.mono (subset_univ _) - -theorem Continuous.continuousWithinAt {f : α → β} {s : Set α} {x : α} (h : Continuous f) : - ContinuousWithinAt f s x := - h.continuousAt.continuousWithinAt - theorem Continuous.comp_continuousOn {g : β → γ} {f : α → β} {s : Set α} (hg : Continuous g) (hf : ContinuousOn f s) : ContinuousOn (g ∘ f) s := hg.continuousOn.comp hf (mapsTo_univ _ _) @@ -839,60 +834,185 @@ theorem ContinuousOn.comp_continuous {g : β → γ} {f : α → β} {s : Set β rw [continuous_iff_continuousOn_univ] at * exact hg.comp hf fun x _ => hs x +theorem ContinuousAt.comp₂_continuousWithinAt {f : β × γ → δ} {g : α → β} {h : α → γ} {x : α} + {s : Set α} (hf : ContinuousAt f (g x, h x)) (hg : ContinuousWithinAt g s x) + (hh : ContinuousWithinAt h s x) : + ContinuousWithinAt (fun x ↦ f (g x, h x)) s x := + ContinuousAt.comp_continuousWithinAt hf (hg.prod_mk_nhds hh) + +theorem ContinuousAt.comp₂_continuousWithinAt_of_eq {f : β × γ → δ} {g : α → β} + {h : α → γ} {x : α} {s : Set α} {y : β × γ} (hf : ContinuousAt f y) + (hg : ContinuousWithinAt g s x) (hh : ContinuousWithinAt h s x) (e : (g x, h x) = y) : + ContinuousWithinAt (fun x ↦ f (g x, h x)) s x := by + rw [← e] at hf + exact hf.comp₂_continuousWithinAt hg hh + + +/-! +### Image +-/ + +theorem ContinuousWithinAt.mem_closure_image {f : α → β} {s : Set α} {x : α} + (h : ContinuousWithinAt f s x) (hx : x ∈ closure s) : f x ∈ closure (f '' s) := + haveI := mem_closure_iff_nhdsWithin_neBot.1 hx + mem_closure_of_tendsto h <| mem_of_superset self_mem_nhdsWithin (subset_preimage_image f s) + +theorem ContinuousWithinAt.mem_closure {f : α → β} {s : Set α} {x : α} {A : Set β} + (h : ContinuousWithinAt f s x) (hx : x ∈ closure s) (hA : MapsTo f s A) : f x ∈ closure A := + closure_mono (image_subset_iff.2 hA) (h.mem_closure_image hx) + +theorem Set.MapsTo.closure_of_continuousWithinAt {f : α → β} {s : Set α} {t : Set β} + (h : MapsTo f s t) (hc : ∀ x ∈ closure s, ContinuousWithinAt f s x) : + MapsTo f (closure s) (closure t) := fun x hx => (hc x hx).mem_closure hx h + +theorem Set.MapsTo.closure_of_continuousOn {f : α → β} {s : Set α} {t : Set β} (h : MapsTo f s t) + (hc : ContinuousOn f (closure s)) : MapsTo f (closure s) (closure t) := + h.closure_of_continuousWithinAt fun x hx => (hc x hx).mono subset_closure + +theorem ContinuousWithinAt.image_closure {f : α → β} {s : Set α} + (hf : ∀ x ∈ closure s, ContinuousWithinAt f s x) : f '' closure s ⊆ closure (f '' s) := + ((mapsTo_image f s).closure_of_continuousWithinAt hf).image_subset + +theorem ContinuousOn.image_closure {f : α → β} {s : Set α} (hf : ContinuousOn f (closure s)) : + f '' closure s ⊆ closure (f '' s) := + ContinuousWithinAt.image_closure fun x hx => (hf x hx).mono subset_closure + +/-! +### Product +-/ + +theorem ContinuousWithinAt.prod_map {f : α → γ} {g : β → δ} {s : Set α} {t : Set β} {x : α} {y : β} + (hf : ContinuousWithinAt f s x) (hg : ContinuousWithinAt g t y) : + ContinuousWithinAt (Prod.map f g) (s ×ˢ t) (x, y) := by + unfold ContinuousWithinAt at * + rw [nhdsWithin_prod_eq, Prod.map, nhds_prod_eq] + exact hf.prod_map hg + +theorem continuousWithinAt_prod_of_discrete_left [DiscreteTopology α] + {f : α × β → γ} {s : Set (α × β)} {x : α × β} : + ContinuousWithinAt f s x ↔ ContinuousWithinAt (f ⟨x.1, ·⟩) {b | (x.1, b) ∈ s} x.2 := by + rw [← x.eta]; simp_rw [ContinuousWithinAt, nhdsWithin, nhds_prod_eq, nhds_discrete, pure_prod, + ← map_inf_principal_preimage]; rfl + +theorem continuousWithinAt_prod_of_discrete_right [DiscreteTopology β] + {f : α × β → γ} {s : Set (α × β)} {x : α × β} : + ContinuousWithinAt f s x ↔ ContinuousWithinAt (f ⟨·, x.2⟩) {a | (a, x.2) ∈ s} x.1 := by + rw [← x.eta]; simp_rw [ContinuousWithinAt, nhdsWithin, nhds_prod_eq, nhds_discrete, prod_pure, + ← map_inf_principal_preimage]; rfl + +theorem continuousAt_prod_of_discrete_left [DiscreteTopology α] {f : α × β → γ} {x : α × β} : + ContinuousAt f x ↔ ContinuousAt (f ⟨x.1, ·⟩) x.2 := by + simp_rw [← continuousWithinAt_univ]; exact continuousWithinAt_prod_of_discrete_left + +theorem continuousAt_prod_of_discrete_right [DiscreteTopology β] {f : α × β → γ} {x : α × β} : + ContinuousAt f x ↔ ContinuousAt (f ⟨·, x.2⟩) x.1 := by + simp_rw [← continuousWithinAt_univ]; exact continuousWithinAt_prod_of_discrete_right + +theorem continuousOn_prod_of_discrete_left [DiscreteTopology α] {f : α × β → γ} {s : Set (α × β)} : + ContinuousOn f s ↔ ∀ a, ContinuousOn (f ⟨a, ·⟩) {b | (a, b) ∈ s} := by + simp_rw [ContinuousOn, Prod.forall, continuousWithinAt_prod_of_discrete_left]; rfl + +theorem continuousOn_prod_of_discrete_right [DiscreteTopology β] {f : α × β → γ} {s : Set (α × β)} : + ContinuousOn f s ↔ ∀ b, ContinuousOn (f ⟨·, b⟩) {a | (a, b) ∈ s} := by + simp_rw [ContinuousOn, Prod.forall, continuousWithinAt_prod_of_discrete_right]; apply forall_swap + +/-- If a function `f a b` is such that `y ↦ f a b` is continuous for all `a`, and `a` lives in a +discrete space, then `f` is continuous, and vice versa. -/ +theorem continuous_prod_of_discrete_left [DiscreteTopology α] {f : α × β → γ} : + Continuous f ↔ ∀ a, Continuous (f ⟨a, ·⟩) := by + simp_rw [continuous_iff_continuousOn_univ]; exact continuousOn_prod_of_discrete_left + +theorem continuous_prod_of_discrete_right [DiscreteTopology β] {f : α × β → γ} : + Continuous f ↔ ∀ b, Continuous (f ⟨·, b⟩) := by + simp_rw [continuous_iff_continuousOn_univ]; exact continuousOn_prod_of_discrete_right + +theorem isOpenMap_prod_of_discrete_left [DiscreteTopology α] {f : α × β → γ} : + IsOpenMap f ↔ ∀ a, IsOpenMap (f ⟨a, ·⟩) := by + simp_rw [isOpenMap_iff_nhds_le, Prod.forall, nhds_prod_eq, nhds_discrete, pure_prod, map_map] + rfl + +theorem isOpenMap_prod_of_discrete_right [DiscreteTopology β] {f : α × β → γ} : + IsOpenMap f ↔ ∀ b, IsOpenMap (f ⟨·, b⟩) := by + simp_rw [isOpenMap_iff_nhds_le, Prod.forall, forall_swap (α := α) (β := β), nhds_prod_eq, + nhds_discrete, prod_pure, map_map]; rfl + +theorem ContinuousOn.prod_map {f : α → γ} {g : β → δ} {s : Set α} {t : Set β} + (hf : ContinuousOn f s) (hg : ContinuousOn g t) : ContinuousOn (Prod.map f g) (s ×ˢ t) := + fun ⟨x, y⟩ ⟨hx, hy⟩ => ContinuousWithinAt.prod_map (hf x hx) (hg y hy) + +theorem ContinuousWithinAt.prod {f : α → β} {g : α → γ} {s : Set α} {x : α} + (hf : ContinuousWithinAt f s x) (hg : ContinuousWithinAt g s x) : + ContinuousWithinAt (fun x => (f x, g x)) s x := + hf.prod_mk_nhds hg + @[fun_prop] -theorem continuousOn_apply {ι : Type*} {π : ι → Type*} [∀ i, TopologicalSpace (π i)] - (i : ι) (s) : ContinuousOn (fun p : ∀ i, π i => p i) s := - Continuous.continuousOn (continuous_apply i) +theorem ContinuousOn.prod {f : α → β} {g : α → γ} {s : Set α} (hf : ContinuousOn f s) + (hg : ContinuousOn g s) : ContinuousOn (fun x => (f x, g x)) s := fun x hx => + ContinuousWithinAt.prod (hf x hx) (hg x hx) -theorem ContinuousWithinAt.preimage_mem_nhdsWithin {f : α → β} {x : α} {s : Set α} {t : Set β} - (h : ContinuousWithinAt f s x) (ht : t ∈ 𝓝 (f x)) : f ⁻¹' t ∈ 𝓝[s] x := - h ht +theorem continuousOn_fst {s : Set (α × β)} : ContinuousOn Prod.fst s := + continuous_fst.continuousOn -theorem Set.LeftInvOn.map_nhdsWithin_eq {f : α → β} {g : β → α} {x : β} {s : Set β} - (h : LeftInvOn f g s) (hx : f (g x) = x) (hf : ContinuousWithinAt f (g '' s) (g x)) - (hg : ContinuousWithinAt g s x) : map g (𝓝[s] x) = 𝓝[g '' s] g x := by - apply le_antisymm - · exact hg.tendsto_nhdsWithin (mapsTo_image _ _) - · have A : g ∘ f =ᶠ[𝓝[g '' s] g x] id := - h.rightInvOn_image.eqOn.eventuallyEq_of_mem self_mem_nhdsWithin - refine le_map_of_right_inverse A ?_ - simpa only [hx] using hf.tendsto_nhdsWithin (h.mapsTo (surjOn_image _ _)) +theorem continuousWithinAt_fst {s : Set (α × β)} {p : α × β} : ContinuousWithinAt Prod.fst s p := + continuous_fst.continuousWithinAt -theorem Function.LeftInverse.map_nhds_eq {f : α → β} {g : β → α} {x : β} - (h : Function.LeftInverse f g) (hf : ContinuousWithinAt f (range g) (g x)) - (hg : ContinuousAt g x) : map g (𝓝 x) = 𝓝[range g] g x := by - simpa only [nhdsWithin_univ, image_univ] using - (h.leftInvOn univ).map_nhdsWithin_eq (h x) (by rwa [image_univ]) hg.continuousWithinAt +@[fun_prop] +theorem ContinuousOn.fst {f : α → β × γ} {s : Set α} (hf : ContinuousOn f s) : + ContinuousOn (fun x => (f x).1) s := + continuous_fst.comp_continuousOn hf -theorem ContinuousWithinAt.preimage_mem_nhdsWithin' {f : α → β} {x : α} {s : Set α} {t : Set β} - (h : ContinuousWithinAt f s x) (ht : t ∈ 𝓝[f '' s] f x) : f ⁻¹' t ∈ 𝓝[s] x := - h.tendsto_nhdsWithin (mapsTo_image _ _) ht +theorem ContinuousWithinAt.fst {f : α → β × γ} {s : Set α} {a : α} (h : ContinuousWithinAt f s a) : + ContinuousWithinAt (fun x => (f x).fst) s a := + continuousAt_fst.comp_continuousWithinAt h -theorem ContinuousWithinAt.preimage_mem_nhdsWithin'' - {f : α → β} {x : α} {y : β} {s t : Set β} - (h : ContinuousWithinAt f (f ⁻¹' s) x) (ht : t ∈ 𝓝[s] y) (hxy : y = f x) : - f ⁻¹' t ∈ 𝓝[f ⁻¹' s] x := by - rw [hxy] at ht - exact h.preimage_mem_nhdsWithin' (nhdsWithin_mono _ (image_preimage_subset f s) ht) +theorem continuousOn_snd {s : Set (α × β)} : ContinuousOn Prod.snd s := + continuous_snd.continuousOn -theorem Filter.EventuallyEq.congr_continuousWithinAt {f g : α → β} {s : Set α} {x : α} - (h : f =ᶠ[𝓝[s] x] g) (hx : f x = g x) : - ContinuousWithinAt f s x ↔ ContinuousWithinAt g s x := by - rw [ContinuousWithinAt, hx, tendsto_congr' h, ContinuousWithinAt] +theorem continuousWithinAt_snd {s : Set (α × β)} {p : α × β} : ContinuousWithinAt Prod.snd s p := + continuous_snd.continuousWithinAt -theorem ContinuousWithinAt.congr_of_eventuallyEq {f f₁ : α → β} {s : Set α} {x : α} - (h : ContinuousWithinAt f s x) (h₁ : f₁ =ᶠ[𝓝[s] x] f) (hx : f₁ x = f x) : - ContinuousWithinAt f₁ s x := - (h₁.congr_continuousWithinAt hx).2 h +@[fun_prop] +theorem ContinuousOn.snd {f : α → β × γ} {s : Set α} (hf : ContinuousOn f s) : + ContinuousOn (fun x => (f x).2) s := + continuous_snd.comp_continuousOn hf -theorem ContinuousWithinAt.congr {f f₁ : α → β} {s : Set α} {x : α} (h : ContinuousWithinAt f s x) - (h₁ : ∀ y ∈ s, f₁ y = f y) (hx : f₁ x = f x) : ContinuousWithinAt f₁ s x := - h.congr_of_eventuallyEq (mem_of_superset self_mem_nhdsWithin h₁) hx +theorem ContinuousWithinAt.snd {f : α → β × γ} {s : Set α} {a : α} (h : ContinuousWithinAt f s a) : + ContinuousWithinAt (fun x => (f x).snd) s a := + continuousAt_snd.comp_continuousWithinAt h -theorem ContinuousWithinAt.congr_mono {f g : α → β} {s s₁ : Set α} {x : α} - (h : ContinuousWithinAt f s x) (h' : EqOn g f s₁) (h₁ : s₁ ⊆ s) (hx : g x = f x) : - ContinuousWithinAt g s₁ x := - (h.mono h₁).congr h' hx +theorem continuousWithinAt_prod_iff {f : α → β × γ} {s : Set α} {x : α} : + ContinuousWithinAt f s x ↔ + ContinuousWithinAt (Prod.fst ∘ f) s x ∧ ContinuousWithinAt (Prod.snd ∘ f) s x := + ⟨fun h => ⟨h.fst, h.snd⟩, fun ⟨h1, h2⟩ => h1.prod h2⟩ + +/-! +### Pi +-/ + +theorem continuousWithinAt_pi {ι : Type*} {π : ι → Type*} [∀ i, TopologicalSpace (π i)] + {f : α → ∀ i, π i} {s : Set α} {x : α} : + ContinuousWithinAt f s x ↔ ∀ i, ContinuousWithinAt (fun y => f y i) s x := + tendsto_pi_nhds + +theorem continuousOn_pi {ι : Type*} {π : ι → Type*} [∀ i, TopologicalSpace (π i)] + {f : α → ∀ i, π i} {s : Set α} : ContinuousOn f s ↔ ∀ i, ContinuousOn (fun y => f y i) s := + ⟨fun h i x hx => tendsto_pi_nhds.1 (h x hx) i, fun h x hx => tendsto_pi_nhds.2 fun i => h i x hx⟩ + +@[fun_prop] +theorem continuousOn_pi' {ι : Type*} {π : ι → Type*} [∀ i, TopologicalSpace (π i)] + {f : α → ∀ i, π i} {s : Set α} (hf : ∀ i, ContinuousOn (fun y => f y i) s) : + ContinuousOn f s := + continuousOn_pi.2 hf + +@[fun_prop] +theorem continuousOn_apply {ι : Type*} {π : ι → Type*} [∀ i, TopologicalSpace (π i)] + (i : ι) (s) : ContinuousOn (fun p : ∀ i, π i => p i) s := + Continuous.continuousOn (continuous_apply i) + + +/-! +### Specific functions +-/ @[fun_prop] theorem continuousOn_const {s : Set α} {c : β} : ContinuousOn (fun _ => c) s := @@ -916,85 +1036,33 @@ protected theorem ContinuousOn.iterate {f : α → α} {s : Set α} (hcont : Con | 0 => continuousOn_id | (n + 1) => (hcont.iterate hmaps n).comp hcont hmaps -theorem continuousOn_open_iff {f : α → β} {s : Set α} (hs : IsOpen s) : - ContinuousOn f s ↔ ∀ t, IsOpen t → IsOpen (s ∩ f ⁻¹' t) := by - rw [continuousOn_iff'] - constructor - · intro h t ht - rcases h t ht with ⟨u, u_open, hu⟩ - rw [inter_comm, hu] - apply IsOpen.inter u_open hs - · intro h t ht - refine ⟨s ∩ f ⁻¹' t, h t ht, ?_⟩ - rw [@inter_comm _ s (f ⁻¹' t), inter_assoc, inter_self] - -theorem ContinuousOn.isOpen_inter_preimage {f : α → β} {s : Set α} {t : Set β} - (hf : ContinuousOn f s) (hs : IsOpen s) (ht : IsOpen t) : IsOpen (s ∩ f ⁻¹' t) := - (continuousOn_open_iff hs).1 hf t ht - -theorem ContinuousOn.isOpen_preimage {f : α → β} {s : Set α} {t : Set β} (h : ContinuousOn f s) - (hs : IsOpen s) (hp : f ⁻¹' t ⊆ s) (ht : IsOpen t) : IsOpen (f ⁻¹' t) := by - convert (continuousOn_open_iff hs).mp h t ht - rw [inter_comm, inter_eq_self_of_subset_left hp] - -theorem ContinuousOn.preimage_isClosed_of_isClosed {f : α → β} {s : Set α} {t : Set β} - (hf : ContinuousOn f s) (hs : IsClosed s) (ht : IsClosed t) : IsClosed (s ∩ f ⁻¹' t) := by - rcases continuousOn_iff_isClosed.1 hf t ht with ⟨u, hu⟩ - rw [inter_comm, hu.2] - apply IsClosed.inter hu.1 hs - -theorem ContinuousOn.preimage_interior_subset_interior_preimage {f : α → β} {s : Set α} {t : Set β} - (hf : ContinuousOn f s) (hs : IsOpen s) : s ∩ f ⁻¹' interior t ⊆ s ∩ interior (f ⁻¹' t) := - calc - s ∩ f ⁻¹' interior t ⊆ interior (s ∩ f ⁻¹' t) := - interior_maximal (inter_subset_inter (Subset.refl _) (preimage_mono interior_subset)) - (hf.isOpen_inter_preimage hs isOpen_interior) - _ = s ∩ interior (f ⁻¹' t) := by rw [interior_inter, hs.interior_eq] - -theorem continuousOn_of_locally_continuousOn {f : α → β} {s : Set α} - (h : ∀ x ∈ s, ∃ t, IsOpen t ∧ x ∈ t ∧ ContinuousOn f (s ∩ t)) : ContinuousOn f s := by - intro x xs - rcases h x xs with ⟨t, open_t, xt, ct⟩ - have := ct x ⟨xs, xt⟩ - rwa [ContinuousWithinAt, ← nhdsWithin_restrict _ xt open_t] at this - -theorem continuousOn_to_generateFrom_iff {β} {s : Set α} {T : Set (Set β)} {f : α → β} : - @ContinuousOn α β _ (.generateFrom T) f s ↔ ∀ x ∈ s, ∀ t ∈ T, f x ∈ t → f ⁻¹' t ∈ 𝓝[s] x := - forall₂_congr fun x _ => by - delta ContinuousWithinAt - simp only [TopologicalSpace.nhds_generateFrom, tendsto_iInf, tendsto_principal, mem_setOf_eq, - and_imp] - exact forall_congr' fun t => forall_swap - --- Porting note: dropped an unneeded assumption -theorem continuousOn_isOpen_of_generateFrom {β : Type*} {s : Set α} {T : Set (Set β)} {f : α → β} - (h : ∀ t ∈ T, IsOpen (s ∩ f ⁻¹' t)) : - @ContinuousOn α β _ (.generateFrom T) f s := - continuousOn_to_generateFrom_iff.2 fun _x hx t ht hxt => mem_nhdsWithin.2 - ⟨_, h t ht, ⟨hx, hxt⟩, fun _y hy => hy.1.2⟩ - -theorem ContinuousWithinAt.prod {f : α → β} {g : α → γ} {s : Set α} {x : α} - (hf : ContinuousWithinAt f s x) (hg : ContinuousWithinAt g s x) : - ContinuousWithinAt (fun x => (f x, g x)) s x := - hf.prod_mk_nhds hg +theorem ContinuousWithinAt.fin_insertNth {n} {π : Fin (n + 1) → Type*} + [∀ i, TopologicalSpace (π i)] (i : Fin (n + 1)) {f : α → π i} {a : α} {s : Set α} + (hf : ContinuousWithinAt f s a) {g : α → ∀ j : Fin n, π (i.succAbove j)} + (hg : ContinuousWithinAt g s a) : ContinuousWithinAt (fun a => i.insertNth (f a) (g a)) s a := + hf.tendsto.fin_insertNth i hg -@[fun_prop] -theorem ContinuousOn.prod {f : α → β} {g : α → γ} {s : Set α} (hf : ContinuousOn f s) - (hg : ContinuousOn g s) : ContinuousOn (fun x => (f x, g x)) s := fun x hx => - ContinuousWithinAt.prod (hf x hx) (hg x hx) +nonrec theorem ContinuousOn.fin_insertNth {n} {π : Fin (n + 1) → Type*} + [∀ i, TopologicalSpace (π i)] (i : Fin (n + 1)) {f : α → π i} {s : Set α} + (hf : ContinuousOn f s) {g : α → ∀ j : Fin n, π (i.succAbove j)} (hg : ContinuousOn g s) : + ContinuousOn (fun a => i.insertNth (f a) (g a)) s := fun a ha => + (hf a ha).fin_insertNth i (hg a ha) -theorem ContinuousAt.comp₂_continuousWithinAt {f : β × γ → δ} {g : α → β} {h : α → γ} {x : α} - {s : Set α} (hf : ContinuousAt f (g x, h x)) (hg : ContinuousWithinAt g s x) - (hh : ContinuousWithinAt h s x) : - ContinuousWithinAt (fun x ↦ f (g x, h x)) s x := - ContinuousAt.comp_continuousWithinAt hf (hg.prod hh) +theorem Set.LeftInvOn.map_nhdsWithin_eq {f : α → β} {g : β → α} {x : β} {s : Set β} + (h : LeftInvOn f g s) (hx : f (g x) = x) (hf : ContinuousWithinAt f (g '' s) (g x)) + (hg : ContinuousWithinAt g s x) : map g (𝓝[s] x) = 𝓝[g '' s] g x := by + apply le_antisymm + · exact hg.tendsto_nhdsWithin (mapsTo_image _ _) + · have A : g ∘ f =ᶠ[𝓝[g '' s] g x] id := + h.rightInvOn_image.eqOn.eventuallyEq_of_mem self_mem_nhdsWithin + refine le_map_of_right_inverse A ?_ + simpa only [hx] using hf.tendsto_nhdsWithin (h.mapsTo (surjOn_image _ _)) -theorem ContinuousAt.comp₂_continuousWithinAt_of_eq {f : β × γ → δ} {g : α → β} - {h : α → γ} {x : α} {s : Set α} {y : β × γ} (hf : ContinuousAt f y) - (hg : ContinuousWithinAt g s x) (hh : ContinuousWithinAt h s x) (e : (g x, h x) = y) : - ContinuousWithinAt (fun x ↦ f (g x, h x)) s x := by - rw [← e] at hf - exact hf.comp₂_continuousWithinAt hg hh +theorem Function.LeftInverse.map_nhds_eq {f : α → β} {g : β → α} {x : β} + (h : Function.LeftInverse f g) (hf : ContinuousWithinAt f (range g) (g x)) + (hg : ContinuousAt g x) : map g (𝓝 x) = 𝓝[range g] g x := by + simpa only [nhdsWithin_univ, image_univ] using + (h.leftInvOn univ).map_nhdsWithin_eq (h x) (by rwa [image_univ]) hg.continuousWithinAt theorem Inducing.continuousWithinAt_iff {f : α → β} {g : β → γ} (hg : Inducing g) {s : Set α} {x : α} : ContinuousWithinAt f s x ↔ ContinuousWithinAt (g ∘ f) s x := by @@ -1024,11 +1092,37 @@ theorem QuotientMap.continuousOn_isOpen_iff {f : α → β} {g : β → γ} (h : simp only [continuousOn_iff_continuous_restrict, (h.restrictPreimage_isOpen hs).continuous_iff] rfl -theorem continuousWithinAt_of_not_mem_closure {f : α → β} {s : Set α} {x : α} (hx : x ∉ closure s) : - ContinuousWithinAt f s x := by - rw [mem_closure_iff_nhdsWithin_neBot, not_neBot] at hx - rw [ContinuousWithinAt, hx] - exact tendsto_bot +theorem IsOpenMap.continuousOn_image_of_leftInvOn {f : α → β} {s : Set α} + (h : IsOpenMap (s.restrict f)) {finv : β → α} (hleft : LeftInvOn finv f s) : + ContinuousOn finv (f '' s) := by + refine continuousOn_iff'.2 fun t ht => ⟨f '' (t ∩ s), ?_, ?_⟩ + · rw [← image_restrict] + exact h _ (ht.preimage continuous_subtype_val) + · rw [inter_eq_self_of_subset_left (image_subset f inter_subset_right), hleft.image_inter'] + +theorem IsOpenMap.continuousOn_range_of_leftInverse {f : α → β} (hf : IsOpenMap f) {finv : β → α} + (hleft : Function.LeftInverse finv f) : ContinuousOn finv (range f) := by + rw [← image_univ] + exact (hf.restrict isOpen_univ).continuousOn_image_of_leftInvOn fun x _ => hleft x + +/-! +### Continuity of piecewise defined functions +-/ + +@[simp] +theorem continuousWithinAt_update_same [DecidableEq α] {f : α → β} {s : Set α} {x : α} {y : β} : + ContinuousWithinAt (update f x y) s x ↔ Tendsto f (𝓝[s \ {x}] x) (𝓝 y) := + calc + ContinuousWithinAt (update f x y) s x ↔ Tendsto (update f x y) (𝓝[s \ {x}] x) (𝓝 y) := by + { rw [← continuousWithinAt_diff_self, ContinuousWithinAt, update_same] } + _ ↔ Tendsto f (𝓝[s \ {x}] x) (𝓝 y) := + tendsto_congr' <| eventually_nhdsWithin_iff.2 <| Eventually.of_forall + fun _ hz => update_noteq hz.2 _ _ + +@[simp] +theorem continuousAt_update_same [DecidableEq α] {f : α → β} {x : α} {y : β} : + ContinuousAt (Function.update f x y) x ↔ Tendsto f (𝓝[≠] x) (𝓝 y) := by + rw [← continuousWithinAt_univ, continuousWithinAt_update_same, compl_eq_univ_diff] theorem ContinuousOn.if' {s : Set α} {p : α → Prop} {f g : α → β} [∀ a, Decidable (p a)] (hpf : ∀ a ∈ s ∩ frontier { a | p a }, @@ -1177,48 +1271,4 @@ theorem continuousOn_piecewise_ite' {s s' t : Set α} {f f' : α → β} [∀ x, theorem continuousOn_piecewise_ite {s s' t : Set α} {f f' : α → β} [∀ x, Decidable (x ∈ t)] (h : ContinuousOn f s) (h' : ContinuousOn f' s') (H : s ∩ frontier t = s' ∩ frontier t) (Heq : EqOn f f' (s ∩ frontier t)) : ContinuousOn (t.piecewise f f') (t.ite s s') := - continuousOn_piecewise_ite' (h.mono inter_subset_left) (h'.mono inter_subset_left) H - Heq - -theorem frontier_inter_open_inter {s t : Set α} (ht : IsOpen t) : - frontier (s ∩ t) ∩ t = frontier s ∩ t := by - simp only [Set.inter_comm _ t, ← Subtype.preimage_coe_eq_preimage_coe_iff, - ht.isOpenMap_subtype_val.preimage_frontier_eq_frontier_preimage continuous_subtype_val, - Subtype.preimage_coe_self_inter] - -theorem continuousOn_fst {s : Set (α × β)} : ContinuousOn Prod.fst s := - continuous_fst.continuousOn - -theorem continuousWithinAt_fst {s : Set (α × β)} {p : α × β} : ContinuousWithinAt Prod.fst s p := - continuous_fst.continuousWithinAt - -@[fun_prop] -theorem ContinuousOn.fst {f : α → β × γ} {s : Set α} (hf : ContinuousOn f s) : - ContinuousOn (fun x => (f x).1) s := - continuous_fst.comp_continuousOn hf - -theorem ContinuousWithinAt.fst {f : α → β × γ} {s : Set α} {a : α} (h : ContinuousWithinAt f s a) : - ContinuousWithinAt (fun x => (f x).fst) s a := - continuousAt_fst.comp_continuousWithinAt h - -theorem continuousOn_snd {s : Set (α × β)} : ContinuousOn Prod.snd s := - continuous_snd.continuousOn - -theorem continuousWithinAt_snd {s : Set (α × β)} {p : α × β} : ContinuousWithinAt Prod.snd s p := - continuous_snd.continuousWithinAt - -@[fun_prop] -theorem ContinuousOn.snd {f : α → β × γ} {s : Set α} (hf : ContinuousOn f s) : - ContinuousOn (fun x => (f x).2) s := - continuous_snd.comp_continuousOn hf - -theorem ContinuousWithinAt.snd {f : α → β × γ} {s : Set α} {a : α} (h : ContinuousWithinAt f s a) : - ContinuousWithinAt (fun x => (f x).snd) s a := - continuousAt_snd.comp_continuousWithinAt h - -theorem continuousWithinAt_prod_iff {f : α → β × γ} {s : Set α} {x : α} : - ContinuousWithinAt f s x ↔ - ContinuousWithinAt (Prod.fst ∘ f) s x ∧ ContinuousWithinAt (Prod.snd ∘ f) s x := - ⟨fun h => ⟨h.fst, h.snd⟩, fun ⟨h1, h2⟩ => h1.prod h2⟩ - -end Pi + continuousOn_piecewise_ite' (h.mono inter_subset_left) (h'.mono inter_subset_left) H Heq From 6c837c78f211be4560cf2b63adb5347e9854b808 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Fri, 18 Oct 2024 12:42:54 +0000 Subject: [PATCH 140/505] chore: delete >1year old deprecations (#17903) --- Mathlib/Algebra/BigOperators/Group/List.lean | 1 - Mathlib/Topology/Algebra/Module/Basic.lean | 5 ----- Mathlib/Topology/Separation.lean | 5 ----- Mathlib/Topology/UniformSpace/CompactConvergence.lean | 5 ----- 4 files changed, 16 deletions(-) diff --git a/Mathlib/Algebra/BigOperators/Group/List.lean b/Mathlib/Algebra/BigOperators/Group/List.lean index 42ca241b7d835..c61a25c867599 100644 --- a/Mathlib/Algebra/BigOperators/Group/List.lean +++ b/Mathlib/Algebra/BigOperators/Group/List.lean @@ -626,7 +626,6 @@ namespace MonoidHom protected theorem map_list_prod (f : M →* N) (l : List M) : f l.prod = (l.map f).prod := map_list_prod f l -attribute [deprecated map_list_prod (since := "2023-01-10")] MonoidHom.map_list_prod attribute [deprecated map_list_sum (since := "2024-05-02")] AddMonoidHom.map_list_sum end MonoidHom diff --git a/Mathlib/Topology/Algebra/Module/Basic.lean b/Mathlib/Topology/Algebra/Module/Basic.lean index 637ef9b981a96..4296c9b917999 100644 --- a/Mathlib/Topology/Algebra/Module/Basic.lean +++ b/Mathlib/Topology/Algebra/Module/Basic.lean @@ -450,11 +450,6 @@ theorem map_smul_of_tower {R S : Type*} [Semiring S] [SMul R M₁] [Module S M f (c • x) = c • f x := LinearMap.CompatibleSMul.map_smul (f : M₁ →ₗ[S] M₂) c x -@[deprecated _root_.map_sum (since := "2023-09-16")] -protected theorem map_sum {ι : Type*} (f : M₁ →SL[σ₁₂] M₂) (s : Finset ι) (g : ι → M₁) : - f (∑ i ∈ s, g i) = ∑ i ∈ s, f (g i) := - map_sum .. - @[simp, norm_cast] theorem coe_coe (f : M₁ →SL[σ₁₂] M₂) : ⇑(f : M₁ →ₛₗ[σ₁₂] M₂) = f := rfl diff --git a/Mathlib/Topology/Separation.lean b/Mathlib/Topology/Separation.lean index 4d39f3cd21001..792e2e67b2544 100644 --- a/Mathlib/Topology/Separation.lean +++ b/Mathlib/Topology/Separation.lean @@ -2122,11 +2122,6 @@ theorem exists_open_between_and_isCompact_closure [LocallyCompactSpace X] [Regul refine ⟨interior L, isOpen_interior, KL, A.trans LU, ?_⟩ exact L_compact.closure_of_subset interior_subset -@[deprecated WeaklyLocallyCompactSpace.locallyCompactSpace (since := "2023-09-03")] -theorem locally_compact_of_compact [T2Space X] [CompactSpace X] : - LocallyCompactSpace X := - inferInstance - end LocallyCompactRegularSpace section T25 diff --git a/Mathlib/Topology/UniformSpace/CompactConvergence.lean b/Mathlib/Topology/UniformSpace/CompactConvergence.lean index f13aaeff613af..c70587a4d1137 100644 --- a/Mathlib/Topology/UniformSpace/CompactConvergence.lean +++ b/Mathlib/Topology/UniformSpace/CompactConvergence.lean @@ -252,11 +252,6 @@ theorem tendsto_iff_tendstoLocallyUniformly [WeaklyLocallyCompactSpace α] : obtain ⟨n, hn₁, hn₂⟩ := exists_compact_mem_nhds x exact ⟨n, hn₂, h n hn₁ V hV⟩ -@[deprecated tendsto_iff_tendstoLocallyUniformly (since := "2023-09-03")] -theorem tendstoLocallyUniformly_of_tendsto [WeaklyLocallyCompactSpace α] (h : Tendsto F p (𝓝 f)) : - TendstoLocallyUniformly (fun i a => F i a) f p := - tendsto_iff_tendstoLocallyUniformly.1 h - section Functorial variable {γ δ : Type*} [TopologicalSpace γ] [UniformSpace δ] From 8f520637f08a283ae0e192c0a218a1b4e67ec8ea Mon Sep 17 00:00:00 2001 From: FR Date: Fri, 18 Oct 2024 13:37:51 +0000 Subject: [PATCH 141/505] chore: redefine `Nat.bitCasesOn` `Nat.binaryRec` (#15571) The new definitions of `Nat.bitCasesOn` and `Nat.binaryRec*` have better runtime performance. Also they have lighter weight dependencies now and may be upstreamed to `Batteries` later. --- Mathlib/Data/Nat/Bits.lean | 133 +++++++++++++++++++--------------- Mathlib/Data/Nat/Bitwise.lean | 10 ++- Mathlib/Data/Nat/Size.lean | 1 - 3 files changed, 79 insertions(+), 65 deletions(-) diff --git a/Mathlib/Data/Nat/Bits.lean b/Mathlib/Data/Nat/Bits.lean index c6e41071fc214..e580011668c94 100644 --- a/Mathlib/Data/Nat/Bits.lean +++ b/Mathlib/Data/Nat/Bits.lean @@ -123,10 +123,24 @@ lemma bit_val (b n) : bit b n = 2 * n + cond b 1 0 := by lemma bit_decomp (n : Nat) : bit (bodd n) (div2 n) = n := (bit_val _ _).trans <| (Nat.add_comm _ _).trans <| bodd_add_div2 _ -/-- For a predicate `C : Nat → Sort*`, if instances can be +theorem shiftRight_one (n) : n >>> 1 = n / 2 := rfl + +@[simp] +theorem bit_decide_mod_two_eq_one_shiftRight_one (n : Nat) : bit (n % 2 = 1) (n >>> 1) = n := by + simp only [bit, shiftRight_one] + cases mod_two_eq_zero_or_one n with | _ h => simpa [h] using Nat.div_add_mod n 2 + +theorem bit_testBit_zero_shiftRight_one (n : Nat) : bit (n.testBit 0) (n >>> 1) = n := by + simp + +/-- For a predicate `motive : Nat → Sort*`, if instances can be constructed for natural numbers of the form `bit b n`, they can be constructed for any given natural number. -/ -def bitCasesOn {C : Nat → Sort u} (n) (h : ∀ b n, C (bit b n)) : C n := bit_decomp n ▸ h _ _ +def bitCasesOn {motive : Nat → Sort u} (n) (h : ∀ b n, motive (bit b n)) : motive n := + -- `1 &&& n != 0` is faster than `n.testBit 0`. This may change when we have faster `testBit`. + let x := h (1 &&& n != 0) (n >>> 1) + -- `congrArg motive _ ▸ x` is defeq to `x` in non-dependent case + congrArg motive n.bit_testBit_zero_shiftRight_one ▸ x lemma bit_zero : bit false 0 = 0 := rfl @@ -158,21 +172,16 @@ lemma binaryRec_decreasing (h : n ≠ 0) : div2 n < n := by rwa [Nat.mul_one] at this /-- A recursion principle for `bit` representations of natural numbers. - For a predicate `C : Nat → Sort*`, if instances can be + For a predicate `motive : Nat → Sort*`, if instances can be constructed for natural numbers of the form `bit b n`, they can be constructed for all natural numbers. -/ -def binaryRec {C : Nat → Sort u} (z : C 0) (f : ∀ b n, C n → C (bit b n)) : ∀ n, C n := - fun n => - if n0 : n = 0 then by - simp only [n0] - exact z - else by - let n' := div2 n - have _x : bit (bodd n) n' = n := by - apply bit_decomp n - rw [← _x] - exact f (bodd n) n' (binaryRec z f n') - decreasing_by exact binaryRec_decreasing n0 +def binaryRec {motive : Nat → Sort u} (z : motive 0) (f : ∀ b n, motive n → motive (bit b n)) + (n : Nat) : motive n := + if n0 : n = 0 then congrArg motive n0 ▸ z + else + let x := f (1 &&& n != 0) (n >>> 1) (binaryRec z f (n >>> 1)) + congrArg motive n.bit_testBit_zero_shiftRight_one ▸ x +decreasing_by exact bitwise_rec_lemma n0 /-- `size n` : Returns the size of a natural number in bits i.e. the length of its binary representation -/ @@ -191,7 +200,8 @@ def ldiff : ℕ → ℕ → ℕ := bitwise fun a b => a && not b @[simp] -lemma binaryRec_zero {C : Nat → Sort u} (z : C 0) (f : ∀ b n, C n → C (bit b n)) : +lemma binaryRec_zero {motive : Nat → Sort u} + (z : motive 0) (f : ∀ b n, motive n → motive (bit b n)) : binaryRec z f 0 = z := by rw [binaryRec] rfl @@ -200,7 +210,8 @@ lemma binaryRec_zero {C : Nat → Sort u} (z : C 0) (f : ∀ b n, C n → C (bit lemma binaryRec_one {C : Nat → Sort u} (z : C 0) (f : ∀ b n, C n → C (bit b n)) : binaryRec z f 1 = f true 0 z := by rw [binaryRec] - simp + simp only [succ_ne_self, ↓reduceDIte, reduceShiftRight, binaryRec_zero] + rfl /-! bitwise ops -/ @@ -249,25 +260,6 @@ lemma testBit_bit_succ (m b n) : testBit (bit b n) (succ m) = testBit n m := by simp only [bodd_eq_one_and_ne_zero] at this exact this -lemma binaryRec_eq {C : Nat → Sort u} {z : C 0} {f : ∀ b n, C n → C (bit b n)} - (h : f false 0 z = z) (b n) : binaryRec z f (bit b n) = f b n (binaryRec z f n) := by - rw [binaryRec] - split_ifs with h' - · generalize binaryRec z f (bit b n) = e - revert e - have bf := bodd_bit b n - have n0 := div2_bit b n - rw [h'] at bf n0 - simp only [bodd_zero, div2_zero] at bf n0 - subst bf n0 - rw [binaryRec_zero] - intros - rw [h, eq_mpr_eq_cast, cast_eq] - · simp only; generalize_proofs h - revert h - rw [bodd_bit, div2_bit] - intros; simp only [eq_mpr_eq_cast, cast_eq] - /-! ### `boddDiv2_eq` and `bodd` -/ @@ -296,28 +288,41 @@ theorem bit_ne_zero (b) {n} (h : n ≠ 0) : bit b n ≠ 0 := by cases b <;> dsimp [bit] <;> omega @[simp] -theorem bitCasesOn_bit {C : ℕ → Sort u} (H : ∀ b n, C (bit b n)) (b : Bool) (n : ℕ) : - bitCasesOn (bit b n) H = H b n := - eq_of_heq <| (eqRec_heq _ _).trans <| by rw [bodd_bit, div2_bit] +theorem bit_div_two (b n) : bit b n / 2 = n := by + rw [bit_val, Nat.add_comm, add_mul_div_left, div_eq_of_lt, Nat.zero_add] + · cases b <;> decide + · decide + +@[simp] +theorem bit_shiftRight_one (b n) : bit b n >>> 1 = n := + bit_div_two b n @[simp] -theorem bitCasesOn_bit0 {C : ℕ → Sort u} (H : ∀ b n, C (bit b n)) (n : ℕ) : +theorem bitCasesOn_bit {motive : ℕ → Sort u} (h : ∀ b n, motive (bit b n)) (b : Bool) (n : ℕ) : + bitCasesOn (bit b n) h = h b n := by + change congrArg motive (bit b n).bit_testBit_zero_shiftRight_one ▸ h _ _ = h b n + generalize congrArg motive (bit b n).bit_testBit_zero_shiftRight_one = e; revert e + rw [testBit_bit_zero, bit_shiftRight_one] + intros; rfl + +@[simp] +theorem bitCasesOn_bit0 {motive : ℕ → Sort u} (H : ∀ b n, motive (bit b n)) (n : ℕ) : bitCasesOn (2 * n) H = H false n := bitCasesOn_bit H false n @[simp] -theorem bitCasesOn_bit1 {C : ℕ → Sort u} (H : ∀ b n, C (bit b n)) (n : ℕ) : +theorem bitCasesOn_bit1 {motive : ℕ → Sort u} (H : ∀ b n, motive (bit b n)) (n : ℕ) : bitCasesOn (2 * n + 1) H = H true n := bitCasesOn_bit H true n -theorem bit_cases_on_injective {C : ℕ → Sort u} : - Function.Injective fun H : ∀ b n, C (bit b n) => fun n => bitCasesOn n H := by +theorem bit_cases_on_injective {motive : ℕ → Sort u} : + Function.Injective fun H : ∀ b n, motive (bit b n) => fun n => bitCasesOn n H := by intro H₁ H₂ h ext b n simpa only [bitCasesOn_bit] using congr_fun h (bit b n) @[simp] -theorem bit_cases_on_inj {C : ℕ → Sort u} (H₁ H₂ : ∀ b n, C (bit b n)) : +theorem bit_cases_on_inj {motive : ℕ → Sort u} (H₁ H₂ : ∀ b n, motive (bit b n)) : ((fun n => bitCasesOn n H₁) = fun n => bitCasesOn n H₂) ↔ H₁ = H₂ := bit_cases_on_injective.eq_iff @@ -340,27 +345,34 @@ The same as `binaryRec_eq`, but that one unfortunately requires `f` to be the identity when appending `false` to `0`. Here, we allow you to explicitly say that that case is not happening, i.e. supplying `n = 0 → b = true`. -/ -theorem binaryRec_eq' {C : ℕ → Sort*} {z : C 0} {f : ∀ b n, C n → C (bit b n)} (b n) - (h : f false 0 z = z ∨ (n = 0 → b = true)) : +theorem binaryRec_eq' {motive : ℕ → Sort*} {z : motive 0} {f : ∀ b n, motive n → motive (bit b n)} + (b n) (h : f false 0 z = z ∨ (n = 0 → b = true)) : binaryRec z f (bit b n) = f b n (binaryRec z f n) := by - rw [binaryRec] - split_ifs with h' - · rcases bit_eq_zero_iff.mp h' with ⟨rfl, rfl⟩ - rw [binaryRec_zero] + by_cases h' : bit b n = 0 + case pos => + obtain ⟨rfl, rfl⟩ := bit_eq_zero_iff.mp h' simp only [imp_false, or_false, eq_self_iff_true, not_true, reduceCtorEq] at h + unfold binaryRec exact h.symm - · dsimp only [] - generalize_proofs e - revert e - rw [bodd_bit, div2_bit] - intros - rfl + case neg => + rw [binaryRec, dif_neg h'] + change congrArg motive (bit b n).bit_testBit_zero_shiftRight_one ▸ f _ _ _ = _ + generalize congrArg motive (bit b n).bit_testBit_zero_shiftRight_one = e; revert e + rw [testBit_bit_zero, bit_shiftRight_one] + intros; rfl + +theorem binaryRec_eq {motive : Nat → Sort u} {z : motive 0} + {f : ∀ b n, motive n → motive (bit b n)} + (h : f false 0 z = z) (b n) : + binaryRec z f (bit b n) = f b n (binaryRec z f n) := + binaryRec_eq' b n (.inl h) /-- The same as `binaryRec`, but the induction step can assume that if `n=0`, the bit being appended is `true`-/ @[elab_as_elim] -def binaryRec' {C : ℕ → Sort*} (z : C 0) (f : ∀ b n, (n = 0 → b = true) → C n → C (bit b n)) : - ∀ n, C n := +def binaryRec' {motive : ℕ → Sort*} (z : motive 0) + (f : ∀ b n, (n = 0 → b = true) → motive n → motive (bit b n)) : + ∀ n, motive n := binaryRec z fun b n ih => if h : n = 0 → b = true then f b n h ih else by @@ -370,8 +382,9 @@ def binaryRec' {C : ℕ → Sort*} (z : C 0) (f : ∀ b n, (n = 0 → b = true) /-- The same as `binaryRec`, but special casing both 0 and 1 as base cases -/ @[elab_as_elim] -def binaryRecFromOne {C : ℕ → Sort*} (z₀ : C 0) (z₁ : C 1) (f : ∀ b n, n ≠ 0 → C n → C (bit b n)) : - ∀ n, C n := +def binaryRecFromOne {motive : ℕ → Sort*} (z₀ : motive 0) (z₁ : motive 1) + (f : ∀ b n, n ≠ 0 → motive n → motive (bit b n)) : + ∀ n, motive n := binaryRec' z₀ fun b n h ih => if h' : n = 0 then by rw [h', h h'] diff --git a/Mathlib/Data/Nat/Bitwise.lean b/Mathlib/Data/Nat/Bitwise.lean index 98bc466c54c8d..2a45dc564a4b8 100644 --- a/Mathlib/Data/Nat/Bitwise.lean +++ b/Mathlib/Data/Nat/Bitwise.lean @@ -69,10 +69,12 @@ lemma bitwise_of_ne_zero {n m : Nat} (hn : n ≠ 0) (hm : m ≠ 0) : theorem binaryRec_of_ne_zero {C : Nat → Sort*} (z : C 0) (f : ∀ b n, C n → C (bit b n)) {n} (h : n ≠ 0) : binaryRec z f n = bit_decomp n ▸ f (bodd n) (div2 n) (binaryRec z f (div2 n)) := by - rw [Eq.rec_eq_cast] - rw [binaryRec] - dsimp only - rw [dif_neg h, eq_mpr_eq_cast] + cases n using bitCasesOn with + | h b n => + rw [binaryRec_eq' _ _ (by right; simpa [bit_eq_zero_iff] using h)] + generalize_proofs h; revert h + rw [bodd_bit, div2_bit] + simp @[simp] lemma bitwise_bit {f : Bool → Bool → Bool} (h : f false false = false := by rfl) (a m b n) : diff --git a/Mathlib/Data/Nat/Size.lean b/Mathlib/Data/Nat/Size.lean index f3f51b3aa4f02..5c7eb08714c3c 100644 --- a/Mathlib/Data/Nat/Size.lean +++ b/Mathlib/Data/Nat/Size.lean @@ -44,7 +44,6 @@ theorem size_bit {b n} (h : bit b n ≠ 0) : size (bit b n) = succ (size n) := b lhs rw [binaryRec] simp [h] - rw [div2_bit] section From e84469bdcf9629d3ca28d01dd2889a5ee3fb5742 Mon Sep 17 00:00:00 2001 From: FR Date: Fri, 18 Oct 2024 13:37:52 +0000 Subject: [PATCH 142/505] =?UTF-8?q?chore:=20do=20not=20use=20`=C2=B7=20?= =?UTF-8?q?=E2=89=88=20=C2=B7`=20in=20`Quotient.eq`=20(#17594)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit When `Setoid` is not an instance, we obviously want to avoid using `· ≈ ·`. Currently mathlib replicates the API around `Quotient.mk` for `Quotient.mk''`, but this is not needed because `Quotient.mk` doesn't use instance parameters now. After adjusting the API we can stop using `Quotient.mk''`. --- Mathlib/Algebra/BigOperators/Group/Finset.lean | 2 +- Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean | 4 ++-- Mathlib/Data/Multiset/Basic.lean | 2 +- Mathlib/Data/Quot.lean | 5 ++++- Mathlib/Data/Setoid/Basic.lean | 7 +++++-- Mathlib/ModelTheory/DirectLimit.lean | 2 +- 6 files changed, 14 insertions(+), 8 deletions(-) diff --git a/Mathlib/Algebra/BigOperators/Group/Finset.lean b/Mathlib/Algebra/BigOperators/Group/Finset.lean index cda549919c99c..ed79af0319d28 100644 --- a/Mathlib/Algebra/BigOperators/Group/Finset.lean +++ b/Mathlib/Algebra/BigOperators/Group/Finset.lean @@ -1618,7 +1618,7 @@ theorem prod_partition (R : Setoid α) [DecidableRel R.r] : /-- If we can partition a product into subsets that cancel out, then the whole product cancels. -/ @[to_additive "If we can partition a sum into subsets that cancel out, then the whole sum cancels."] theorem prod_cancels_of_partition_cancels (R : Setoid α) [DecidableRel R.r] - (h : ∀ x ∈ s, ∏ a ∈ s.filter fun y => y ≈ x, f a = 1) : ∏ x ∈ s, f x = 1 := by + (h : ∀ x ∈ s, ∏ a ∈ s.filter fun y => R y x, f a = 1) : ∏ x ∈ s, f x = 1 := by rw [prod_partition R, ← Finset.prod_eq_one] intro xbar xbar_in_s obtain ⟨x, x_in_s, rfl⟩ := mem_image.mp xbar_in_s diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean index d0976e216dd81..e4838de2f5dd0 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean @@ -1169,7 +1169,7 @@ lemma add_of_equiv {P Q : Fin 3 → R} (h : P ≈ Q) : W'.add P Q = W'.dblXYZ P lemma add_smul_of_equiv {P Q : Fin 3 → R} (h : P ≈ Q) {u v : R} (hu : IsUnit u) (hv : IsUnit v) : W'.add (u • P) (v • Q) = u ^ 4 • W'.add P Q := by have smul : P ≈ Q ↔ u • P ≈ v • Q := by - erw [← Quotient.eq, ← Quotient.eq, smul_eq P hu, smul_eq Q hv] + erw [← Quotient.eq_iff_equiv, ← Quotient.eq_iff_equiv, smul_eq P hu, smul_eq Q hv] rfl rw [add_of_equiv <| smul.mp h, dblXYZ_smul, add_of_equiv h] @@ -1185,7 +1185,7 @@ lemma add_of_not_equiv {P Q : Fin 3 → R} (h : ¬P ≈ Q) : W'.add P Q = W'.add lemma add_smul_of_not_equiv {P Q : Fin 3 → R} (h : ¬P ≈ Q) {u v : R} (hu : IsUnit u) (hv : IsUnit v) : W'.add (u • P) (v • Q) = (u * v) ^ 2 • W'.add P Q := by have smul : P ≈ Q ↔ u • P ≈ v • Q := by - erw [← Quotient.eq, ← Quotient.eq, smul_eq P hu, smul_eq Q hv] + erw [← Quotient.eq_iff_equiv, ← Quotient.eq_iff_equiv, smul_eq P hu, smul_eq Q hv] rfl rw [add_of_not_equiv <| h.comp smul.mpr, addXYZ_smul, add_of_not_equiv h] diff --git a/Mathlib/Data/Multiset/Basic.lean b/Mathlib/Data/Multiset/Basic.lean index 8f15a338625d1..e04ef7fa7a1dc 100644 --- a/Mathlib/Data/Multiset/Basic.lean +++ b/Mathlib/Data/Multiset/Basic.lean @@ -69,7 +69,7 @@ instance [DecidableEq α] (l₁ l₂ : List α) : Decidable (l₁ ≈ l₂) := -- Porting note: `Quotient.recOnSubsingleton₂ s₁ s₂` was in parens which broke elaboration instance decidableEq [DecidableEq α] : DecidableEq (Multiset α) - | s₁, s₂ => Quotient.recOnSubsingleton₂ s₁ s₂ fun _ _ => decidable_of_iff' _ Quotient.eq + | s₁, s₂ => Quotient.recOnSubsingleton₂ s₁ s₂ fun _ _ => decidable_of_iff' _ Quotient.eq_iff_equiv /-- defines a size for a multiset by referring to the size of the underlying list -/ protected diff --git a/Mathlib/Data/Quot.lean b/Mathlib/Data/Quot.lean index 8f524efb7730b..a11fa73582072 100644 --- a/Mathlib/Data/Quot.lean +++ b/Mathlib/Data/Quot.lean @@ -274,9 +274,12 @@ theorem Quot.eq {α : Type*} {r : α → α → Prop} {x y : α} : ⟨Quot.eqvGen_exact, Quot.eqvGen_sound⟩ @[simp] -theorem Quotient.eq {r : Setoid α} {x y : α} : Quotient.mk r x = ⟦y⟧ ↔ x ≈ y := +theorem Quotient.eq {r : Setoid α} {x y : α} : Quotient.mk r x = ⟦y⟧ ↔ r x y := ⟨Quotient.exact, Quotient.sound⟩ +theorem Quotient.eq_iff_equiv {r : Setoid α} {x y : α} : Quotient.mk r x = ⟦y⟧ ↔ x ≈ y := + Quotient.eq + theorem Quotient.forall {α : Sort*} {s : Setoid α} {p : Quotient s → Prop} : (∀ a, p a) ↔ ∀ a : α, p ⟦a⟧ := ⟨fun h _ ↦ h _, fun h a ↦ a.ind h⟩ diff --git a/Mathlib/Data/Setoid/Basic.lean b/Mathlib/Data/Setoid/Basic.lean index 63241cbb3689d..c34439b36e1e1 100644 --- a/Mathlib/Data/Setoid/Basic.lean +++ b/Mathlib/Data/Setoid/Basic.lean @@ -221,9 +221,12 @@ lemma sInf_equiv {S : Set (Setoid α)} {x y : α} : letI := sInf S x ≈ y ↔ ∀ s ∈ S, s.Rel x y := Iff.rfl +lemma sInf_iff {S : Set (Setoid α)} {x y : α} : + sInf S x y ↔ ∀ s ∈ S, s x y := Iff.rfl + lemma quotient_mk_sInf_eq {S : Set (Setoid α)} {x y : α} : - Quotient.mk (sInf S) x = Quotient.mk (sInf S) y ↔ ∀ s ∈ S, s.Rel x y := by - simp [sInf_equiv] + Quotient.mk (sInf S) x = Quotient.mk (sInf S) y ↔ ∀ s ∈ S, s x y := by + simp [sInf_iff] /-- The map induced between quotients by a setoid inequality. -/ def map_of_le {s t : Setoid α} (h : s ≤ t) : Quotient s → Quotient t := diff --git a/Mathlib/ModelTheory/DirectLimit.lean b/Mathlib/ModelTheory/DirectLimit.lean index 295d49fb6ebf9..bc5cfde3f62f1 100644 --- a/Mathlib/ModelTheory/DirectLimit.lean +++ b/Mathlib/ModelTheory/DirectLimit.lean @@ -349,7 +349,7 @@ def lift (g : ∀ i, G i ↪[L] P) (Hg : ∀ i j hij x, g j (f i j hij x) = g i rw [← Quotient.out_eq x, ← Quotient.out_eq y, Quotient.lift_mk, Quotient.lift_mk] at xy obtain ⟨i, hx, hy⟩ := directed_of (· ≤ ·) x.out.1 y.out.1 rw [← Hg x.out.1 i hx, ← Hg y.out.1 i hy] at xy - rw [← Quotient.out_eq x, ← Quotient.out_eq y, Quotient.eq, equiv_iff G f hx hy] + rw [← Quotient.out_eq x, ← Quotient.out_eq y, Quotient.eq_iff_equiv, equiv_iff G f hx hy] exact (g i).injective xy map_fun' F x := by obtain ⟨i, y, rfl⟩ := exists_quotient_mk'_sigma_mk'_eq G f x From 344c0bb0e582a25235a93bb531130cb0433c6392 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 18 Oct 2024 14:18:26 +0000 Subject: [PATCH 143/505] feat: a subgroup is `1` iff it's trivial (#17906) From LeanCamCombi --- Mathlib/Algebra/Group/Subgroup/Pointwise.lean | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/Mathlib/Algebra/Group/Subgroup/Pointwise.lean b/Mathlib/Algebra/Group/Subgroup/Pointwise.lean index 95c62df222f42..f42601fa4cc09 100644 --- a/Mathlib/Algebra/Group/Subgroup/Pointwise.lean +++ b/Mathlib/Algebra/Group/Subgroup/Pointwise.lean @@ -41,6 +41,10 @@ lemma smul_coe_set [Group G] [SetLike S G] [SubgroupClass S G] {s : S} {a : G} ( a • (s : Set G) = s := by ext; simp [Set.mem_smul_set_iff_inv_smul_mem, mul_mem_cancel_left, ha] +@[norm_cast, to_additive] +lemma coe_set_eq_one [Group G] {s : Subgroup G} : (s : Set G) = 1 ↔ s = ⊥ := + (SetLike.ext'_iff.trans (by rfl)).symm + @[to_additive (attr := simp)] lemma op_smul_coe_set [Group G] [SetLike S G] [SubgroupClass S G] {s : S} {a : G} (ha : a ∈ s) : MulOpposite.op a • (s : Set G) = s := by From 2a3ae36b046fab48ee07482c6b66d7d5f1e6eff3 Mon Sep 17 00:00:00 2001 From: Scott Carnahan <128885296+ScottCarnahan@users.noreply.github.com> Date: Fri, 18 Oct 2024 14:51:42 +0000 Subject: [PATCH 144/505] feat (RingTheory/HahnSeries/Summable) : transport summable families along Equiv (#16872) This PR defines a new summable family given a summable family of Hahn series and an equivalence of parametrizing types. We show that the new family has the same formal sum as the old one. --- Mathlib/RingTheory/HahnSeries/Summable.lean | 25 ++++++++++++++++----- 1 file changed, 20 insertions(+), 5 deletions(-) diff --git a/Mathlib/RingTheory/HahnSeries/Summable.lean b/Mathlib/RingTheory/HahnSeries/Summable.lean index 4405c005d34e1..45a21761c48c3 100644 --- a/Mathlib/RingTheory/HahnSeries/Summable.lean +++ b/Mathlib/RingTheory/HahnSeries/Summable.lean @@ -42,7 +42,7 @@ open Pointwise noncomputable section -variable {Γ : Type*} {R : Type*} +variable {Γ Γ' R V α β : Type*} namespace HahnSeries @@ -61,13 +61,11 @@ theorem isPWO_iUnion_support_powers [LinearOrderedCancelAddCommMonoid Γ] [Ring section -variable (Γ) (R) [PartialOrder Γ] [AddCommMonoid R] - /-- A family of Hahn series whose formal coefficient-wise sum is a Hahn series. For each coefficient of the sum to be well-defined, we require that only finitely many series are nonzero at any given coefficient. For the formal sum to be a Hahn series, we require that the union of the supports of the constituent series is partially well-ordered. -/ -structure SummableFamily (α : Type*) where +structure SummableFamily (Γ) (R) [PartialOrder Γ] [AddCommMonoid R] (α : Type*) where /-- A parametrized family of Hahn series. -/ toFun : α → HahnSeries Γ R isPWO_iUnion_support' : Set.IsPWO (⋃ a : α, (toFun a).support) @@ -79,7 +77,7 @@ namespace SummableFamily section AddCommMonoid -variable [PartialOrder Γ] [AddCommMonoid R] {α : Type*} +variable [PartialOrder Γ] [AddCommMonoid R] instance : FunLike (SummableFamily Γ R α) α (HahnSeries Γ R) where coe := toFun @@ -180,6 +178,23 @@ theorem hsum_add {s t : SummableFamily Γ R α} : (s + t).hsum = s.hsum + t.hsum simp only [hsum_coeff, add_coeff, add_apply] exact finsum_add_distrib (s.finite_co_support _) (t.finite_co_support _) +/-- A summable family induced by an equivalence of the parametrizing type. -/ +@[simps] +def Equiv (e : α ≃ β) (s : SummableFamily Γ R α) : SummableFamily Γ R β where + toFun b := s (e.symm b) + isPWO_iUnion_support' := by + refine Set.IsPWO.mono s.isPWO_iUnion_support fun g => ?_ + simp only [Set.mem_iUnion, mem_support, ne_eq, forall_exists_index] + exact fun b hg => Exists.intro (e.symm b) hg + finite_co_support' g := + (Equiv.set_finite_iff e.subtypeEquivOfSubtype').mp <| s.finite_co_support' g + +@[simp] +theorem hsum_equiv (e : α ≃ β) (s : SummableFamily Γ R α) : (Equiv e s).hsum = s.hsum := by + ext g + simp only [hsum_coeff, Equiv_toFun] + exact finsum_eq_of_bijective e.symm (Equiv.bijective e.symm) fun x => rfl + end AddCommMonoid section AddCommGroup From c289ac29f16b5d6733857a95139dc57fa8ebc646 Mon Sep 17 00:00:00 2001 From: Jz Pan Date: Fri, 18 Oct 2024 14:51:43 +0000 Subject: [PATCH 145/505] feat(RingTheory/Algebraic): add some results on `Transcendental` (#17867) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit - `transcendental_iff`, `transcendental_iff_ker_eq_bot`: similar to the API of `AlgebraicIndependent` - `IsAlgebraic.of_aeval`, `Transcendental.aeval`: behavior of transcendental over `aeval` - `Polynomial.transcendental`, `Polynomial.transcendental_X`: some specific polynomials are transcendental over its base ring - `Transcendental.linearIndependent_sub_inv`: if `E / F` is a field extension, `x` is an element of `E` transcendental over `F`, then `{(x - a)⁻¹ | a : F}` is linearly independent over `F` - `transcendental_algebraMap_iff`: negation of `isAlgebraic_algebraMap_iff` --- Mathlib/RingTheory/Algebraic.lean | 69 +++++++++++++++++++++++++++++++ 1 file changed, 69 insertions(+) diff --git a/Mathlib/RingTheory/Algebraic.lean b/Mathlib/RingTheory/Algebraic.lean index 8685279579238..9942e6c347371 100644 --- a/Mathlib/RingTheory/Algebraic.lean +++ b/Mathlib/RingTheory/Algebraic.lean @@ -40,6 +40,63 @@ theorem is_transcendental_of_subsingleton [Subsingleton R] (x : A) : Transcenden variable {R} +/-- An element `x` is transcendental over `R` if and only if for any polynomial `p`, +`Polynomial.aeval x p = 0` implies `p = 0`. This is similar to `algebraicIndependent_iff`. -/ +theorem transcendental_iff {x : A} : + Transcendental R x ↔ ∀ p : R[X], aeval x p = 0 → p = 0 := by + rw [Transcendental, IsAlgebraic, not_exists] + congr! 1; tauto + +variable (R) in +theorem Polynomial.transcendental_X : Transcendental R (X (R := R)) := by + simp [transcendental_iff] + +theorem IsAlgebraic.of_aeval {r : A} (f : R[X]) (hf : f.natDegree ≠ 0) + (hf' : f.leadingCoeff ∈ nonZeroDivisors R) (H : IsAlgebraic R (aeval r f)) : + IsAlgebraic R r := by + obtain ⟨p, h1, h2⟩ := H + have : (p.comp f).coeff (p.natDegree * f.natDegree) ≠ 0 := fun h ↦ h1 <| by + rwa [coeff_comp_degree_mul_degree hf, + mul_right_mem_nonZeroDivisors_eq_zero_iff (pow_mem hf' _), + leadingCoeff_eq_zero] at h + exact ⟨p.comp f, fun h ↦ this (by simp [h]), by rwa [aeval_comp]⟩ + +theorem Transcendental.aeval {r : A} (H : Transcendental R r) (f : R[X]) (hf : f.natDegree ≠ 0) + (hf' : f.leadingCoeff ∈ nonZeroDivisors R) : + Transcendental R (aeval r f) := fun h ↦ H (h.of_aeval f hf hf') + +theorem Polynomial.transcendental (f : R[X]) (hf : f.natDegree ≠ 0) + (hf' : f.leadingCoeff ∈ nonZeroDivisors R) : + Transcendental R f := by + simpa using (transcendental_X R).aeval f hf hf' + +/-- If `E / F` is a field extension, `x` is an element of `E` transcendental over `F`, +then `{(x - a)⁻¹ | a : F}` is linearly independent over `F`. -/ +theorem Transcendental.linearIndependent_sub_inv + {F E : Type*} [Field F] [Field E] [Algebra F E] {x : E} (H : Transcendental F x) : + LinearIndependent F fun a ↦ (x - algebraMap F E a)⁻¹ := by + rw [transcendental_iff] at H + refine linearIndependent_iff'.2 fun s m hm i hi ↦ ?_ + have hnz (a : F) : x - algebraMap F E a ≠ 0 := fun h ↦ + X_sub_C_ne_zero a <| H (.X - .C a) (by simp [h]) + let b := s.prod fun j ↦ x - algebraMap F E j + have h1 : ∀ i ∈ s, m i • (b * (x - algebraMap F E i)⁻¹) = + m i • (s.erase i).prod fun j ↦ x - algebraMap F E j := fun i hi ↦ by + simp_rw [b, ← s.prod_erase_mul _ hi, mul_inv_cancel_right₀ (hnz i)] + replace hm := congr(b * $(hm)) + simp_rw [mul_zero, Finset.mul_sum, mul_smul_comm, Finset.sum_congr rfl h1] at hm + let p : Polynomial F := s.sum fun i ↦ .C (m i) * (s.erase i).prod fun j ↦ .X - .C j + replace hm := congr(Polynomial.aeval i $(H p (by simp_rw [← hm, p, map_sum, map_mul, map_prod, + map_sub, aeval_X, aeval_C, Algebra.smul_def]))) + have h2 : ∀ j ∈ s.erase i, m j * ((s.erase j).prod fun x ↦ i - x) = 0 := fun j hj ↦ by + have := Finset.mem_erase_of_ne_of_mem (Finset.ne_of_mem_erase hj).symm hi + simp_rw [← (s.erase j).prod_erase_mul _ this, sub_self, mul_zero] + simp_rw [map_zero, p, map_sum, map_mul, map_prod, map_sub, aeval_X, + aeval_C, Algebra.id.map_eq_self, ← s.sum_erase_add _ hi, + Finset.sum_eq_zero h2, zero_add] at hm + exact eq_zero_of_ne_zero_of_mul_right_eq_zero (Finset.prod_ne_zero_iff.2 fun j hj ↦ + sub_ne_zero.2 (Finset.ne_of_mem_erase hj).symm) hm + /-- A subalgebra is algebraic if all its elements are algebraic. -/ nonrec def Subalgebra.IsAlgebraic (S : Subalgebra R A) : Prop := @@ -86,10 +143,18 @@ theorem isAlgebraic_iff_not_injective {x : A} : IsAlgebraic R x ↔ ¬Function.Injective (Polynomial.aeval x : R[X] →ₐ[R] A) := by simp only [IsAlgebraic, injective_iff_map_eq_zero, not_forall, and_comm, exists_prop] +/-- An element `x` is transcendental over `R` if and only if the map `Polynomial.aeval x` +is injective. This is similar to `algebraicIndependent_iff_injective_aeval`. -/ theorem transcendental_iff_injective {x : A} : Transcendental R x ↔ Function.Injective (Polynomial.aeval x : R[X] →ₐ[R] A) := isAlgebraic_iff_not_injective.not_left +/-- An element `x` is transcendental over `R` if and only if the kernel of the ring homomorphism +`Polynomial.aeval x` is the zero ideal. This is similar to `algebraicIndependent_iff_ker_eq_bot`. -/ +theorem transcendental_iff_ker_eq_bot {x : A} : + Transcendental R x ↔ RingHom.ker (aeval (R := R) x) = ⊥ := by + rw [transcendental_iff_injective, RingHom.injective_iff_ker_eq_bot] + end section zero_ne_one @@ -174,6 +239,10 @@ theorem isAlgebraic_algebraMap_iff {a : S} (h : Function.Injective (algebraMap S IsAlgebraic R (algebraMap S A a) ↔ IsAlgebraic R a := isAlgebraic_algHom_iff (IsScalarTower.toAlgHom R S A) h +theorem transcendental_algebraMap_iff {a : S} (h : Function.Injective (algebraMap S A)) : + Transcendental R (algebraMap S A a) ↔ Transcendental R a := by + simp_rw [Transcendental, isAlgebraic_algebraMap_iff h] + theorem IsAlgebraic.of_pow {r : A} {n : ℕ} (hn : 0 < n) (ht : IsAlgebraic R (r ^ n)) : IsAlgebraic R r := by obtain ⟨p, p_nonzero, hp⟩ := ht From 66fe0b7279d0edf044d905b5a4068aa929b78e7b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 18 Oct 2024 14:51:44 +0000 Subject: [PATCH 146/505] doc(Simps): deduplicate clause (#17916) Both clauses were different before #8296 --- Mathlib/Tactic/Simps/Basic.lean | 3 +-- 1 file changed, 1 insertion(+), 2 deletions(-) diff --git a/Mathlib/Tactic/Simps/Basic.lean b/Mathlib/Tactic/Simps/Basic.lean index 1caba0e4825ab..3b28a506c4d61 100644 --- a/Mathlib/Tactic/Simps/Basic.lean +++ b/Mathlib/Tactic/Simps/Basic.lean @@ -317,8 +317,7 @@ Some common uses: This will generate `foo_apply` lemmas for each declaration `foo`. * If you prefer `coe_foo` lemmas that state equalities between functions, use `initialize_simps_projections MulHom (toFun → coe, as_prefix coe)` - In this case you have to use `@[simps (config := .asFn)]` or equivalently - `@[simps (config := .asFn)]` whenever you call `@[simps]`. + In this case you have to use `@[simps (config := .asFn)]` whenever you call `@[simps]`. * You can also initialize to use both, in which case you have to choose which one to use by default, by using either of the following ``` From 9e3fa1e4f5106b00d9b078c99479dae53c6adbb7 Mon Sep 17 00:00:00 2001 From: Jz Pan Date: Fri, 18 Oct 2024 15:39:36 +0000 Subject: [PATCH 147/505] feat(FieldTheory/Adjoin): add some results of `extendScalars` (#15238) - `extendScalars_{self|top|inf|sup}` for intermediate fields and subfields - also add `restrictScalars_{inf|sup}` and `bot_toSubfield` - add a missing helper lemma `toSubfield_injective` --- Mathlib/FieldTheory/Adjoin.lean | 82 ++++++++++++++++++- .../FieldTheory/IntermediateField/Basic.lean | 15 ++-- 2 files changed, 89 insertions(+), 8 deletions(-) diff --git a/Mathlib/FieldTheory/Adjoin.lean b/Mathlib/FieldTheory/Adjoin.lean index aa48dc241889b..f5504ce434121 100644 --- a/Mathlib/FieldTheory/Adjoin.lean +++ b/Mathlib/FieldTheory/Adjoin.lean @@ -114,6 +114,9 @@ theorem mem_bot {x : E} : x ∈ (⊥ : IntermediateField F E) ↔ x ∈ Set.rang @[simp] theorem bot_toSubalgebra : (⊥ : IntermediateField F E).toSubalgebra = ⊥ := rfl +theorem bot_toSubfield : (⊥ : IntermediateField F E).toSubfield = (algebraMap F E).fieldRange := + rfl + @[simp] theorem coe_top : ↑(⊤ : IntermediateField F E) = (Set.univ : Set E) := rfl @@ -267,16 +270,31 @@ def topEquiv : (⊤ : IntermediateField F E) ≃ₐ[F] E := -- Porting note: this theorem is now generated by the `@[simps!]` above. +section RestrictScalars + @[simp] theorem restrictScalars_bot_eq_self (K : IntermediateField F E) : (⊥ : IntermediateField K E).restrictScalars _ = K := SetLike.coe_injective Subtype.range_coe +variable {K : Type*} [Field K] [Algebra K E] [Algebra K F] [IsScalarTower K F E] + @[simp] -theorem restrictScalars_top {K : Type*} [Field K] [Algebra K E] [Algebra K F] - [IsScalarTower K F E] : (⊤ : IntermediateField F E).restrictScalars K = ⊤ := +theorem restrictScalars_top : (⊤ : IntermediateField F E).restrictScalars K = ⊤ := rfl +variable (K) +variable (L L' : IntermediateField F E) + +theorem restrictScalars_sup : + L.restrictScalars K ⊔ L'.restrictScalars K = (L ⊔ L').restrictScalars K := + toSubfield_injective (by simp) + +theorem restrictScalars_inf : + L.restrictScalars K ⊓ L'.restrictScalars K = (L ⊓ L').restrictScalars K := rfl + +end RestrictScalars + variable {K : Type*} [Field K] [Algebra F K] @[simp] @@ -1466,3 +1484,63 @@ theorem comap_map (f : L →ₐ[K] L') (S : IntermediateField K L) : (S.map f).c SetLike.coe_injective (Set.preimage_image_eq _ f.injective) end IntermediateField + +section ExtendScalars + +variable {K : Type*} [Field K] {L : Type*} [Field L] [Algebra K L] + +namespace Subfield + +variable (F : Subfield L) + +@[simp] +theorem extendScalars_self : extendScalars (le_refl F) = ⊥ := by + ext x + rw [mem_extendScalars, IntermediateField.mem_bot] + refine ⟨fun h ↦ ⟨⟨x, h⟩, rfl⟩, ?_⟩ + rintro ⟨y, rfl⟩ + exact y.2 + +@[simp] +theorem extendScalars_top : extendScalars (le_top : F ≤ ⊤) = ⊤ := + IntermediateField.toSubfield_injective (by simp) + +variable {F} +variable {E E' : Subfield L} (h : F ≤ E) (h' : F ≤ E') + +theorem extendScalars_sup : + extendScalars h ⊔ extendScalars h' = extendScalars (le_sup_of_le_left h : F ≤ E ⊔ E') := + ((extendScalars.orderIso F).map_sup ⟨_, h⟩ ⟨_, h'⟩).symm + +theorem extendScalars_inf : extendScalars h ⊓ extendScalars h' = extendScalars (le_inf h h') := + ((extendScalars.orderIso F).map_inf ⟨_, h⟩ ⟨_, h'⟩).symm + +end Subfield + +namespace IntermediateField + +variable (F : IntermediateField K L) + +@[simp] +theorem extendScalars_self : extendScalars (le_refl F) = ⊥ := + restrictScalars_injective K (by simp) + +@[simp] +theorem extendScalars_top : extendScalars (le_top : F ≤ ⊤) = ⊤ := + restrictScalars_injective K (by simp) + +variable {F} +variable {E E' : IntermediateField K L} (h : F ≤ E) (h' : F ≤ E') + +theorem extendScalars_sup : + extendScalars h ⊔ extendScalars h' = extendScalars (le_sup_of_le_left h : F ≤ E ⊔ E') := + ((extendScalars.orderIso F).map_sup ⟨_, h⟩ ⟨_, h'⟩).symm + +theorem extendScalars_inf : extendScalars h ⊓ extendScalars h' = extendScalars (le_inf h h') := + ((extendScalars.orderIso F).map_inf ⟨_, h⟩ ⟨_, h'⟩).symm + +end IntermediateField + +end ExtendScalars + +set_option linter.style.longFile 1700 diff --git a/Mathlib/FieldTheory/IntermediateField/Basic.lean b/Mathlib/FieldTheory/IntermediateField/Basic.lean index 1c807c5364473..54dbf4d6b374d 100644 --- a/Mathlib/FieldTheory/IntermediateField/Basic.lean +++ b/Mathlib/FieldTheory/IntermediateField/Basic.lean @@ -494,14 +494,17 @@ theorem coe_inclusion {E F : IntermediateField K L} (hEF : E ≤ F) (e : E) : variable {S} -theorem toSubalgebra_injective : - Function.Injective (IntermediateField.toSubalgebra : IntermediateField K L → _) := by - intro S S' h +theorem toSubalgebra_injective : Function.Injective (toSubalgebra : IntermediateField K L → _) := by + intro _ _ h + ext + simp_rw [← mem_toSubalgebra, h] + +theorem toSubfield_injective : Function.Injective (toSubfield : IntermediateField K L → _) := by + intro _ _ h ext - rw [← mem_toSubalgebra, ← mem_toSubalgebra, h] + simp_rw [← mem_toSubfield, h] -theorem map_injective (f : L →ₐ[K] L') : - Function.Injective (map f) := by +theorem map_injective (f : L →ₐ[K] L') : Function.Injective (map f) := by intro _ _ h rwa [← toSubalgebra_injective.eq_iff, toSubalgebra_map, toSubalgebra_map, (Subalgebra.map_injective f.injective).eq_iff, toSubalgebra_injective.eq_iff] at h From c5b97e2f7eecede13acd99994d4be2239dbd2ea3 Mon Sep 17 00:00:00 2001 From: Yongle Hu Date: Fri, 18 Oct 2024 15:39:37 +0000 Subject: [PATCH 148/505] feat(NumberTheory/NumberField/Basic): lemmas for extension of ring of integers (#17784) Add some lemmas and `instance`s for extension of number fields or ring of integers. Co-authored-by: Yongle Hu @mbkybky Co-authored-by: Jiedong Jiang @jjdishere Co-authored-by: Hu Yongle <2065545849@qq.com> --- Mathlib/NumberTheory/NumberField/Basic.lean | 64 ++++++++++++++++++++- 1 file changed, 63 insertions(+), 1 deletion(-) diff --git a/Mathlib/NumberTheory/NumberField/Basic.lean b/Mathlib/NumberTheory/NumberField/Basic.lean index cd142c881c370..64c8f95606e20 100644 --- a/Mathlib/NumberTheory/NumberField/Basic.lean +++ b/Mathlib/NumberTheory/NumberField/Basic.lean @@ -59,6 +59,23 @@ protected theorem isAlgebraic [NumberField K] : Algebra.IsAlgebraic ℚ K := instance [NumberField K] [NumberField L] [Algebra K L] : FiniteDimensional K L := Module.Finite.of_restrictScalars_finite ℚ K L +/-- A finite extension of a number field is a number field. -/ +theorem of_module_finite [NumberField K] [Algebra K L] [Module.Finite K L] : NumberField L where + to_charZero := charZero_of_injective_algebraMap (algebraMap K L).injective + to_finiteDimensional := + letI := charZero_of_injective_algebraMap (algebraMap K L).injective + Module.Finite.trans K L + +variable {K} {L} in +instance of_intermediateField [NumberField K] [NumberField L] [Algebra K L] + (E : IntermediateField K L) : NumberField E := + of_module_finite K E + +theorem of_tower [NumberField K] [NumberField L] [Algebra K L] (E : Type*) [Field E] + [Algebra K E] [Algebra E L] [IsScalarTower K E L] : NumberField E := + letI := Module.Finite.left K E L + of_module_finite K E + /-- The ring of integers (or number ring) corresponding to a number field is the integral closure of ℤ in the number field. @@ -89,7 +106,7 @@ instance : Nontrivial (𝓞 K) := inferInstanceAs (Nontrivial (integralClosure _ _)) instance {L : Type*} [Ring L] [Algebra K L] : Algebra (𝓞 K) L := inferInstanceAs (Algebra (integralClosure _ _) L) -instance {L : Type*} [Ring L] [Algebra K L] : IsScalarTower (𝓞 K) K L := +instance {L : Type*} [Ring L] [Algebra K L] : IsScalarTower (𝓞 K) K L := inferInstanceAs (IsScalarTower (integralClosure _ _) K L) variable {K} @@ -222,6 +239,9 @@ instance [NumberField K] : IsFractionRing (𝓞 K) K := instance : IsIntegralClosure (𝓞 K) ℤ K := integralClosure.isIntegralClosure _ _ +instance : Algebra.IsIntegral ℤ (𝓞 K) := + IsIntegralClosure.isIntegral_algebra ℤ K + instance [NumberField K] : IsIntegrallyClosed (𝓞 K) := integralClosure.isIntegrallyClosedOfFiniteExtension ℚ @@ -283,6 +303,48 @@ def restrict_monoidHom [MulOneClass M] (f : M →* K) (h : ∀ x, IsIntegral ℤ map_one' := by simp only [restrict, map_one, mk_one] map_mul' x y := by simp only [restrict, map_mul, mk_mul_mk _] +section extension + +variable (K L : Type*) [Field K] [Field L] [Algebra K L] + +instance : IsScalarTower (𝓞 K) (𝓞 L) L := + IsScalarTower.of_algebraMap_eq' rfl + +instance : IsIntegralClosure (𝓞 L) (𝓞 K) L := + IsIntegralClosure.tower_top (R := ℤ) + +/-- The ring of integers of `L` is isomorphic to any integral closure of `𝓞 K` in `L` -/ +protected noncomputable def algEquiv (R : Type*) [CommRing R] [Algebra (𝓞 K) R] [Algebra R L] + [IsScalarTower (𝓞 K) R L] [IsIntegralClosure R (𝓞 K) L] : 𝓞 L ≃ₐ[𝓞 K] R := + (IsIntegralClosure.equiv (𝓞 K) R L _).symm + +/-- Any extension between ring of integers is integral. -/ +instance extension_algebra_isIntegral : Algebra.IsIntegral (𝓞 K) (𝓞 L) := + IsIntegralClosure.isIntegral_algebra (𝓞 K) L + +/-- Any extension between ring of integers of number fields is noetherian. -/ +instance extension_isNoetherian [NumberField K] [NumberField L] : IsNoetherian (𝓞 K) (𝓞 L) := + IsIntegralClosure.isNoetherian (𝓞 K) K L (𝓞 L) + +/-- The kernel of the algebraMap between ring of integers is `⊥`. -/ +theorem ker_algebraMap_eq_bot : RingHom.ker (algebraMap (𝓞 K) (𝓞 L)) = ⊥ := + (RingHom.ker_eq_bot_iff_eq_zero (algebraMap (𝓞 K) (𝓞 L))).mpr <| fun x hx => by + have h : (algebraMap K L) x = (algebraMap (𝓞 K) (𝓞 L)) x := rfl + simp only [hx, map_zero, map_eq_zero, RingOfIntegers.coe_eq_zero_iff] at h + exact h + +/-- The algebraMap between ring of integers is injective. -/ +theorem algebraMap.injective : Function.Injective (algebraMap (𝓞 K) (𝓞 L)) := + (RingHom.injective_iff_ker_eq_bot (algebraMap (𝓞 K) (𝓞 L))).mpr (ker_algebraMap_eq_bot K L) + +instance : NoZeroSMulDivisors (𝓞 K) (𝓞 L) := + NoZeroSMulDivisors.of_algebraMap_injective (algebraMap.injective K L) + +instance : NoZeroSMulDivisors (𝓞 K) L := + NoZeroSMulDivisors.trans (𝓞 K) (𝓞 L) L + +end extension + end RingOfIntegers variable [NumberField K] From 8b31f1fbde480a5c530264eeea00d051d4c586c5 Mon Sep 17 00:00:00 2001 From: Yakov Pechersky Date: Fri, 18 Oct 2024 15:48:17 +0000 Subject: [PATCH 149/505] feat(GroupTheory/ArchimedeanDensely): discrete iff not densely ordered (#16618) Stronger statements than previously proven `Or`s --- Mathlib/GroupTheory/ArchimedeanDensely.lean | 108 ++++++++++++++++---- 1 file changed, 86 insertions(+), 22 deletions(-) diff --git a/Mathlib/GroupTheory/ArchimedeanDensely.lean b/Mathlib/GroupTheory/ArchimedeanDensely.lean index 8b00a52cd33de..cf344500028b4 100644 --- a/Mathlib/GroupTheory/ArchimedeanDensely.lean +++ b/Mathlib/GroupTheory/ArchimedeanDensely.lean @@ -191,6 +191,21 @@ lemma LinearOrderedAddCommGroup.discrete_or_denselyOrdered (G : Type*) · simp [hz.left] · simpa [lt_sub_iff_add_lt'] using hz.right +/-- Any linearly ordered archimedean additive group is either isomorphic (and order-isomorphic) +to the integers, or is densely ordered, exclusively. -/ +lemma LinearOrderedAddCommGroup.discrete_iff_not_denselyOrdered (G : Type*) + [LinearOrderedAddCommGroup G] [Archimedean G] : + Nonempty (G ≃+o ℤ) ↔ ¬ DenselyOrdered G := by + suffices ∀ (_ : G ≃+o ℤ), ¬ DenselyOrdered G by + rcases LinearOrderedAddCommGroup.discrete_or_denselyOrdered G with ⟨⟨h⟩⟩|h + · simpa [this h] using ⟨h⟩ + · simp only [h, not_true_eq_false, iff_false, not_nonempty_iff] + exact ⟨fun H ↦ (this H) h⟩ + intro e H + rw [denselyOrdered_iff_of_orderIsoClass e] at H + obtain ⟨_, _⟩ := exists_between (one_pos (α := ℤ)) + linarith + variable (G) in /-- Any linearly ordered mul-archimedean group is either isomorphic (and order-isomorphic) to the multiplicative integers, or is densely ordered. -/ @@ -201,31 +216,28 @@ lemma LinearOrderedCommGroup.discrete_or_denselyOrdered : rintro ⟨f, hf⟩ exact ⟨AddEquiv.toMultiplicative' f, hf⟩ -/-- Any nontrivial (has other than 0 and 1) linearly ordered mul-archimedean group with zero is -either isomorphic (and order-isomorphic) to `ℤₘ₀`, or is densely ordered. -/ -lemma LinearOrderedCommGroupWithZero.discrete_or_denselyOrdered (G : Type*) - [LinearOrderedCommGroupWithZero G] [Nontrivial Gˣ] [MulArchimedean G] : - Nonempty (G ≃*o ℤₘ₀) ∨ DenselyOrdered G := by - classical - refine (LinearOrderedCommGroup.discrete_or_denselyOrdered Gˣ).imp ?_ ?_ - · intro ⟨f⟩ - refine ⟨OrderMonoidIso.trans - ⟨WithZero.withZeroUnitsEquiv.symm, ?_⟩ ⟨f.withZero, ?_⟩⟩ - · intro - simp only [WithZero.withZeroUnitsEquiv, MulEquiv.symm_mk, - MulEquiv.toEquiv_eq_coe, Equiv.toFun_as_coe, EquivLike.coe_coe, MulEquiv.coe_mk, - Equiv.coe_fn_symm_mk ] - split_ifs <;> - simp_all [← Units.val_le_val] - · intro a b - induction a <;> induction b <;> - simp [MulEquiv.withZero] +variable (G) in +/-- Any linearly ordered mul-archimedean group is either isomorphic (and order-isomorphic) +to the multiplicative integers, or is densely ordered, exclusively. -/ +@[to_additive existing] +lemma LinearOrderedCommGroup.discrete_iff_not_denselyOrdered : + Nonempty (G ≃*o Multiplicative ℤ) ↔ ¬ DenselyOrdered G := by + let e : G ≃o Additive G := OrderIso.refl G + rw [denselyOrdered_iff_of_orderIsoClass e, + ← LinearOrderedAddCommGroup.discrete_iff_not_denselyOrdered (Additive G)] + refine Nonempty.congr ?_ ?_ <;> intro f + · exact ⟨MulEquiv.toAdditive' f, by simp⟩ + · exact ⟨MulEquiv.toAdditive'.symm f, by simp⟩ + +lemma denselyOrdered_units_iff {G₀ : Type*} [LinearOrderedCommGroupWithZero G₀] [Nontrivial G₀ˣ] : + DenselyOrdered G₀ˣ ↔ DenselyOrdered G₀ := by + constructor · intro H refine ⟨fun x y h ↦ ?_⟩ rcases (zero_le' (a := x)).eq_or_lt with rfl|hx - · lift y to Gˣ using h.ne'.isUnit - obtain ⟨z, hz⟩ := exists_ne (1 : Gˣ) - refine ⟨(y * |z|ₘ⁻¹ : Gˣ), ?_, ?_⟩ + · lift y to G₀ˣ using h.ne'.isUnit + obtain ⟨z, hz⟩ := exists_ne (1 : G₀ˣ) + refine ⟨(y * |z|ₘ⁻¹ : G₀ˣ), ?_, ?_⟩ · simp [zero_lt_iff] · rw [Units.val_lt_val] simp [hz] @@ -233,3 +245,55 @@ lemma LinearOrderedCommGroupWithZero.discrete_or_denselyOrdered (G : Type*) (by simp [← Units.val_lt_val, h]) refine ⟨z, ?_, ?_⟩ <;> simpa [← Units.val_lt_val] + · intro H + refine ⟨fun x y h ↦ ?_⟩ + obtain ⟨z, hz⟩ := exists_between (Units.val_lt_val.mpr h) + rcases (zero_le' (a := z)).eq_or_lt with rfl|hz' + · simp at hz + refine ⟨Units.mk0 z hz'.ne', ?_⟩ + simp [← Units.val_lt_val, hz] + +/-- Any nontrivial (has other than 0 and 1) linearly ordered mul-archimedean group with zero is +either isomorphic (and order-isomorphic) to `ℤₘ₀`, or is densely ordered. -/ +lemma LinearOrderedCommGroupWithZero.discrete_or_denselyOrdered (G : Type*) + [LinearOrderedCommGroupWithZero G] [Nontrivial Gˣ] [MulArchimedean G] : + Nonempty (G ≃*o WithZero (Multiplicative ℤ)) ∨ DenselyOrdered G := by + classical + rw [← denselyOrdered_units_iff] + refine (LinearOrderedCommGroup.discrete_or_denselyOrdered Gˣ).imp_left ?_ + intro ⟨f⟩ + refine ⟨OrderMonoidIso.trans + ⟨WithZero.withZeroUnitsEquiv.symm, ?_⟩ ⟨f.withZero, ?_⟩⟩ + · intro + simp only [WithZero.withZeroUnitsEquiv, MulEquiv.symm_mk, + MulEquiv.toEquiv_eq_coe, Equiv.toFun_as_coe, EquivLike.coe_coe, MulEquiv.coe_mk, + Equiv.coe_fn_symm_mk ] + split_ifs <;> + simp_all [← Units.val_le_val] + · intro a b + induction a <;> induction b <;> + simp [MulEquiv.withZero] + +open WithZero in +/-- Any nontrivial (has other than 0 and 1) linearly ordered mul-archimedean group with zero is +either isomorphic (and order-isomorphic) to `ℤₘ₀`, or is densely ordered, exclusively -/ +lemma LinearOrderedCommGroupWithZero.discrete_iff_not_denselyOrdered (G : Type*) + [LinearOrderedCommGroupWithZero G] [Nontrivial Gˣ] [MulArchimedean G] : + Nonempty (G ≃*o WithZero (Multiplicative ℤ)) ↔ ¬ DenselyOrdered G := by + rw [← denselyOrdered_units_iff, + ← LinearOrderedCommGroup.discrete_iff_not_denselyOrdered] + refine Nonempty.congr ?_ ?_ <;> intro f + · refine ⟨MulEquiv.unzero (withZeroUnitsEquiv.trans f), ?_⟩ + intros + simp only [MulEquiv.unzero, withZeroUnitsEquiv, MulEquiv.trans_apply, + MulEquiv.coe_mk, Equiv.coe_fn_mk, recZeroCoe_coe, OrderMonoidIso.coe_mulEquiv, + MulEquiv.symm_trans_apply, MulEquiv.symm_mk, Equiv.coe_fn_symm_mk, map_eq_zero, coe_ne_zero, + ↓reduceDIte, unzero_coe, MulEquiv.toEquiv_eq_coe, Equiv.toFun_as_coe, EquivLike.coe_coe] + rw [← Units.val_le_val, ← map_le_map_iff f, ← coe_le_coe, coe_unzero, coe_unzero] + · refine ⟨withZeroUnitsEquiv.symm.trans (MulEquiv.withZero f), ?_⟩ + intros + simp only [withZeroUnitsEquiv, MulEquiv.symm_mk, MulEquiv.withZero, + MulEquiv.toMonoidHom_eq_coe, MulEquiv.toEquiv_eq_coe, Equiv.toFun_as_coe, EquivLike.coe_coe, + MulEquiv.trans_apply, MulEquiv.coe_mk, Equiv.coe_fn_symm_mk, Equiv.coe_fn_mk] + split_ifs <;> + simp_all [← Units.val_le_val] From 94bb3a8e89c75b20472ec5a37df0f225cb08e003 Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Fri, 18 Oct 2024 15:58:51 +0000 Subject: [PATCH 150/505] feat(CategoryTheory/Monoidal): add `Mon_Class` type class (#17185) [Zulip discussion](https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/Algebras.20in.20monoidal.20categories.2C.20bundled.20or.20semi-bundled.3F) --- Mathlib/CategoryTheory/Monoidal/Bimon_.lean | 41 +++++++++++- Mathlib/CategoryTheory/Monoidal/Comon_.lean | 67 ++++++++++++++++++- Mathlib/CategoryTheory/Monoidal/Hopf_.lean | 35 +++++++++- Mathlib/CategoryTheory/Monoidal/Mon_.lean | 72 ++++++++++++++++++++- 4 files changed, 208 insertions(+), 7 deletions(-) diff --git a/Mathlib/CategoryTheory/Monoidal/Bimon_.lean b/Mathlib/CategoryTheory/Monoidal/Bimon_.lean index 109f528c80123..e195b65504704 100644 --- a/Mathlib/CategoryTheory/Monoidal/Bimon_.lean +++ b/Mathlib/CategoryTheory/Monoidal/Bimon_.lean @@ -29,7 +29,46 @@ universe v₁ v₂ u₁ u₂ u open CategoryTheory MonoidalCategory -variable (C : Type u₁) [Category.{v₁} C] [MonoidalCategory.{v₁} C] [BraidedCategory C] +variable {C : Type u₁} [Category.{v₁} C] [MonoidalCategory.{v₁} C] [BraidedCategory C] + +open scoped Mon_Class Comon_Class + +/-- +A bimonoid object in a braided category `C` is a object that is simultaneously monoid and comonoid +objects, and structure morphisms of them satisfy appropriate consistency conditions. +-/ +class Bimon_Class (M : C) extends Mon_Class M, Comon_Class M where + /- For the names of the conditions below, the unprimed names are reserved for the version where + the argument `M` is explicit. -/ + mul_comul' : μ[M] ≫ Δ[M] = (Δ[M] ⊗ Δ[M]) ≫ tensor_μ M M M M ≫ (μ[M] ⊗ μ[M]) := by aesop_cat + one_comul' : η[M] ≫ Δ[M] = η[M ⊗ M] := by aesop_cat + mul_counit' : μ[M] ≫ ε[M] = ε[M ⊗ M] := by aesop_cat + one_counit' : η[M] ≫ ε[M] = 𝟙 (𝟙_ C) := by aesop_cat + +namespace Bimon_Class + +/- The simp attribute is reserved for the unprimed versions. -/ +attribute [reassoc] mul_comul' one_comul' mul_counit' one_counit' + +variable (M : C) [Bimon_Class M] + +@[reassoc (attr := simp)] +theorem mul_comul (M : C) [Bimon_Class M] : + μ[M] ≫ Δ[M] = (Δ[M] ⊗ Δ[M]) ≫ tensor_μ M M M M ≫ (μ[M] ⊗ μ[M]) := + mul_comul' + +@[reassoc (attr := simp)] +theorem one_comul (M : C) [Bimon_Class M] : η[M] ≫ Δ[M] = η[M ⊗ M] := one_comul' + +@[reassoc (attr := simp)] +theorem mul_counit (M : C) [Bimon_Class M] : μ[M] ≫ ε[M] = ε[M ⊗ M] := mul_counit' + +@[reassoc (attr := simp)] +theorem one_counit (M : C) [Bimon_Class M] : η[M] ≫ ε[M] = 𝟙 (𝟙_ C) := one_counit' + +end Bimon_Class + +variable (C) /-- A bimonoid object in a braided category `C` is a comonoid object in the (monoidal) diff --git a/Mathlib/CategoryTheory/Monoidal/Comon_.lean b/Mathlib/CategoryTheory/Monoidal/Comon_.lean index 498a20b561c79..ff63ce4bbb0eb 100644 --- a/Mathlib/CategoryTheory/Monoidal/Comon_.lean +++ b/Mathlib/CategoryTheory/Monoidal/Comon_.lean @@ -29,7 +29,49 @@ universe v₁ v₂ u₁ u₂ u open CategoryTheory MonoidalCategory -variable (C : Type u₁) [Category.{v₁} C] [MonoidalCategory.{v₁} C] +variable {C : Type u₁} [Category.{v₁} C] [MonoidalCategory.{v₁} C] + +/-- A comonoid object internal to a monoidal category. + +When the monoidal category is preadditive, this is also sometimes called a "coalgebra object". +-/ +class Comon_Class (X : C) where + /-- The counit morphism of a comonoid object. -/ + counit : X ⟶ 𝟙_ C + /-- The comultiplication morphism of a comonoid object. -/ + comul : X ⟶ X ⊗ X + /- For the names of the conditions below, the unprimed names are reserved for the version where + the argument `X` is explicit. -/ + counit_comul' : comul ≫ counit ▷ X = (λ_ X).inv := by aesop_cat + comul_counit' : comul ≫ X ◁ counit = (ρ_ X).inv := by aesop_cat + comul_assoc' : comul ≫ X ◁ comul = comul ≫ (comul ▷ X) ≫ (α_ X X X).hom := by aesop_cat + +namespace Comon_Class + +@[inherit_doc] scoped notation "Δ" => Comon_Class.comul +@[inherit_doc] scoped notation "Δ["M"]" => Comon_Class.comul (X := M) +@[inherit_doc] scoped notation "ε" => Comon_Class.counit +@[inherit_doc] scoped notation "ε["M"]" => Comon_Class.counit (X := M) + +/- The simp attribute is reserved for the unprimed versions. -/ +attribute [reassoc] counit_comul' comul_counit' comul_assoc' + +@[reassoc (attr := simp)] +theorem counit_comul (X : C) [Comon_Class X] : Δ ≫ ε ▷ X = (λ_ X).inv := counit_comul' + +@[reassoc (attr := simp)] +theorem comul_counit (X : C) [Comon_Class X] : Δ ≫ X ◁ ε = (ρ_ X).inv := comul_counit' + +@[reassoc (attr := simp)] +theorem comul_assoc (X : C) [Comon_Class X] : + Δ ≫ X ◁ Δ = Δ ≫ Δ ▷ X ≫ (α_ X X X).hom := + comul_assoc' + +end Comon_Class + +open scoped Comon_Class + +variable (C) /-- A comonoid object internal to a monoidal category. @@ -52,6 +94,24 @@ attribute [reassoc (attr := simp)] Comon_.comul_assoc namespace Comon_ +variable {C} + +/-- Construct an object of `Comon_ C` from an object `X : C` and `Comon_Class X` instance. -/ +@[simps] +def mk' (X : C) [Comon_Class X] : Comon_ C where + X := X + counit := ε + comul := Δ + +instance {M : Comon_ C} : Comon_Class M.X where + counit := M.counit + comul := M.comul + counit_comul' := M.counit_comul + comul_counit' := M.comul_counit + comul_assoc' := M.comul_assoc + +variable (C) + /-- The trivial comonoid object. We later show this is terminal in `Comon_ C`. -/ @[simps] @@ -251,6 +311,9 @@ variable [BraidedCategory C] theorem tensorObj_X (A B : Comon_ C) : (A ⊗ B).X = A.X ⊗ B.X := rfl +instance (A B : C) [Comon_Class A] [Comon_Class B] : Comon_Class (A ⊗ B) := + inferInstanceAs <| Comon_Class (Comon_.mk' A ⊗ Comon_.mk' B).X + theorem tensorObj_counit (A B : Comon_ C) : (A ⊗ B).counit = (A.counit ⊗ B.counit) ≫ (λ_ _).hom := rfl @@ -281,7 +344,7 @@ theorem tensorObj_comul (A B : Comon_ C) : /-- The forgetful functor from `Comon_ C` to `C` is monoidal when `C` is braided monoidal. -/ def forgetMonoidal : MonoidalFunctor (Comon_ C) C := { forget C with - ε := 𝟙 _ + «ε» := 𝟙 _ μ := fun _ _ => 𝟙 _ } @[simp] theorem forgetMonoidal_toFunctor : (forgetMonoidal C).toFunctor = forget C := rfl diff --git a/Mathlib/CategoryTheory/Monoidal/Hopf_.lean b/Mathlib/CategoryTheory/Monoidal/Hopf_.lean index 4bf50784e7cc7..e9e5220ef1ddc 100644 --- a/Mathlib/CategoryTheory/Monoidal/Hopf_.lean +++ b/Mathlib/CategoryTheory/Monoidal/Hopf_.lean @@ -22,7 +22,40 @@ universe v₁ v₂ u₁ u₂ u open CategoryTheory MonoidalCategory -variable (C : Type u₁) [Category.{v₁} C] [MonoidalCategory.{v₁} C] [BraidedCategory C] +variable {C : Type u₁} [Category.{v₁} C] [MonoidalCategory.{v₁} C] [BraidedCategory C] + +open scoped Mon_Class Comon_Class + +/-- +A Hopf monoid in a braided category `C` is a bimonoid object in `C` equipped with an antipode. +-/ +class Hopf_Class (X : C) extends Bimon_Class X where + /-- The antipode is an endomorphism of the underlying object of the Hopf monoid. -/ + antipode : X ⟶ X + /- For the names of the conditions below, the unprimed names are reserved for the version where + the argument `X` is explicit. -/ + antipode_left' : Δ ≫ antipode ▷ X ≫ μ = ε ≫ η := by aesop_cat + antipode_right' : Δ ≫ X ◁ antipode ≫ μ = ε ≫ η := by aesop_cat + +namespace Hopf_Class + +@[inherit_doc] scoped notation "𝒮" => Hopf_Class.antipode +@[inherit_doc] scoped notation "𝒮["X"]" => Hopf_Class.antipode (X := X) + +/- The simp attribute is reserved for the unprimed versions. -/ +attribute [reassoc] antipode_left' antipode_right' + +/-- The object is provided as an explicit argument. -/ +@[reassoc (attr := simp)] +theorem antipode_left (X : C) [Hopf_Class X] : Δ ≫ 𝒮 ▷ X ≫ μ = ε ≫ η := antipode_left' + +/-- The object is provided as an explicit argument. -/ +@[reassoc (attr := simp)] +theorem antipode_right (X : C) [Hopf_Class X] : Δ ≫ X ◁ 𝒮 ≫ μ = ε ≫ η := antipode_right' + +end Hopf_Class + +variable (C) /-- A Hopf monoid in a braided category `C` is a bimonoid object in `C` equipped with an antipode. diff --git a/Mathlib/CategoryTheory/Monoidal/Mon_.lean b/Mathlib/CategoryTheory/Monoidal/Mon_.lean index fc0cec7c2faf7..c04c4f4d55c07 100644 --- a/Mathlib/CategoryTheory/Monoidal/Mon_.lean +++ b/Mathlib/CategoryTheory/Monoidal/Mon_.lean @@ -23,7 +23,51 @@ universe v₁ v₂ u₁ u₂ u open CategoryTheory MonoidalCategory -variable (C : Type u₁) [Category.{v₁} C] [MonoidalCategory.{v₁} C] +variable {C : Type u₁} [Category.{v₁} C] [MonoidalCategory.{v₁} C] + +/-- A monoid object internal to a monoidal category. + +When the monoidal category is preadditive, this is also sometimes called an "algebra object". +-/ +class Mon_Class (X : C) where + /-- The unit morphism of a monoid object. -/ + one : 𝟙_ C ⟶ X + /-- The multiplication morphism of a monoid object. -/ + mul : X ⊗ X ⟶ X + /- For the names of the conditions below, the unprimed names are reserved for the version where + the argument `X` is explicit. -/ + one_mul' : one ▷ X ≫ mul = (λ_ X).hom := by aesop_cat + mul_one' : X ◁ one ≫ mul = (ρ_ X).hom := by aesop_cat + -- Obviously there is some flexibility stating this axiom. + -- This one has left- and right-hand sides matching the statement of `Monoid.mul_assoc`, + -- and chooses to place the associator on the right-hand side. + -- The heuristic is that unitors and associators "don't have much weight". + mul_assoc' : (mul ▷ X) ≫ mul = (α_ X X X).hom ≫ (X ◁ mul) ≫ mul := by aesop_cat + +namespace Mon_Class + +@[inherit_doc] scoped notation "μ" => Mon_Class.mul +@[inherit_doc] scoped notation "μ["M"]" => Mon_Class.mul (X := M) +@[inherit_doc] scoped notation "η" => Mon_Class.one +@[inherit_doc] scoped notation "η["M"]" => Mon_Class.one (X := M) + +/- The simp attribute is reserved for the unprimed versions. -/ +attribute [reassoc] one_mul' mul_one' mul_assoc' + +@[reassoc (attr := simp)] +theorem one_mul (X : C) [Mon_Class X] : η ▷ X ≫ μ = (λ_ X).hom := one_mul' + +@[reassoc (attr := simp)] +theorem mul_one (X : C) [Mon_Class X] : X ◁ η ≫ μ = (ρ_ X).hom := mul_one' + +@[reassoc (attr := simp)] +theorem mul_assoc (X : C) [Mon_Class X] : μ ▷ X ≫ μ = (α_ X X X).hom ≫ X ◁ μ ≫ μ := mul_assoc' + +end Mon_Class + +open scoped Mon_Class + +variable (C) /-- A monoid object internal to a monoidal category. @@ -50,6 +94,24 @@ attribute [reassoc (attr := simp)] Mon_.mul_assoc namespace Mon_ +variable {C} + +/-- Construct an object of `Mon_ C` from an object `X : C` and `Mon_Class X` instance. -/ +@[simps] +def mk' (X : C) [Mon_Class X] : Mon_ C where + X := X + one := η + mul := μ + +instance {M : Mon_ C} : Mon_Class M.X where + one := M.one + mul := M.mul + one_mul' := M.one_mul + mul_one' := M.mul_one + mul_assoc' := M.mul_assoc + +variable (C) + /-- The trivial monoid object. We later show this is initial in `Mon_ C`. -/ @[simps] @@ -235,7 +297,7 @@ def monToLaxMonoidal : Mon_ C ⥤ LaxMonoidalFunctor (Discrete PUnit.{u + 1}) C { obj := fun _ => A.X map := fun _ => 𝟙 _ ε := A.one - μ := fun _ _ => A.mul } + «μ» := fun _ _ => A.mul } map f := { app := fun _ => f.hom } @@ -496,13 +558,17 @@ theorem tensor_mul (M N : Mon_ C) : (M ⊗ N).mul = instance monMonoidal : MonoidalCategory (Mon_ C) where tensorHom_def := by intros; ext; simp [tensorHom_def] +@[simps!] +instance {M N : C} [Mon_Class M] [Mon_Class N] : Mon_Class (M ⊗ N) := + inferInstanceAs <| Mon_Class (Mon_.mk' M ⊗ Mon_.mk' N).X + variable (C) /-- The forgetful functor from `Mon_ C` to `C` is monoidal when `C` is braided monoidal. -/ def forgetMonoidal : MonoidalFunctor (Mon_ C) C := { forget C with ε := 𝟙 _ - μ := fun _ _ => 𝟙 _ } + «μ» := fun _ _ => 𝟙 _ } @[simp] theorem forgetMonoidal_toFunctor : (forgetMonoidal C).toFunctor = forget C := rfl @[simp] theorem forgetMonoidal_ε : (forgetMonoidal C).ε = 𝟙 (𝟙_ C) := rfl From c6a02ffa70f93156c098b09f6b8d612ed1eb8a52 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Fri, 18 Oct 2024 16:20:05 +0000 Subject: [PATCH 151/505] chore: more `erw`s that can become `rw`s (#17923) Automatically found by the `erw` linter from #17638. --- Mathlib/CategoryTheory/Limits/Shapes/Reflexive.lean | 2 +- Mathlib/CategoryTheory/Shift/Basic.lean | 2 +- Mathlib/Condensed/Discrete/Colimit.lean | 4 ++-- Mathlib/Geometry/RingedSpace/OpenImmersion.lean | 4 ++-- Mathlib/Topology/Order/ScottTopology.lean | 4 ++-- 5 files changed, 8 insertions(+), 8 deletions(-) diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Reflexive.lean b/Mathlib/CategoryTheory/Limits/Shapes/Reflexive.lean index cb15f1505053c..abcc9d116cda0 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Reflexive.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Reflexive.lean @@ -514,7 +514,7 @@ lemma mk_π {X : C} (π : F.obj zero ⟶ X) (h : F.map left ≫ π = F.map right (mk π h).π = π := rfl lemma condition (G : ReflexiveCofork F) : F.map left ≫ G.π = F.map right ≫ G.π := by - erw [Cocone.w G left, Cocone.w G right] + rw [Cocone.w G left, Cocone.w G right] @[simp] lemma app_one_eq_π (G : ReflexiveCofork F) : G.ι.app zero = G.π := rfl diff --git a/Mathlib/CategoryTheory/Shift/Basic.lean b/Mathlib/CategoryTheory/Shift/Basic.lean index 7acb86cbee29d..889c3bf069f82 100644 --- a/Mathlib/CategoryTheory/Shift/Basic.lean +++ b/Mathlib/CategoryTheory/Shift/Basic.lean @@ -519,7 +519,7 @@ lemma shiftFunctorCompIsoId_add'_inv_app : dsimp [shiftFunctorCompIsoId] simp only [Functor.map_comp, Category.assoc] congr 1 - erw [← NatTrans.naturality] + rw [← NatTrans.naturality] dsimp rw [← cancel_mono ((shiftFunctorAdd' C p' p 0 hp).inv.app X), Iso.hom_inv_id_app, Category.assoc, Category.assoc, Category.assoc, Category.assoc, diff --git a/Mathlib/Condensed/Discrete/Colimit.lean b/Mathlib/Condensed/Discrete/Colimit.lean index 75f3c6ee7dfeb..9e6930f3aa849 100644 --- a/Mathlib/Condensed/Discrete/Colimit.lean +++ b/Mathlib/Condensed/Discrete/Colimit.lean @@ -276,7 +276,7 @@ lemma isoLocallyConstantOfIsColimit_inv (X : Profinite.{u}ᵒᵖ ⥤ Type (u+1)) simp only [types_comp_apply, isoFinYoneda_inv_app, counitApp_app] apply presheaf_ext.{u, u+1} (X := X) (Y := X) (f := f) intro x - erw [incl_of_counitAppApp] + rw [incl_of_counitAppApp] simp only [counitAppAppImage, CompHausLike.coe_of] letI : Fintype (fiber.{u, u+1} f x) := Fintype.ofInjective (sigmaIncl.{u, u+1} f x).1 Subtype.val_injective @@ -551,7 +551,7 @@ lemma isoLocallyConstantOfIsColimit_inv (X : LightProfinite.{u}ᵒᵖ ⥤ Type u simp only [types_comp_apply, isoFinYoneda_inv_app, counitApp_app] apply presheaf_ext.{u, u} (X := X) (Y := X) (f := f) intro x - erw [incl_of_counitAppApp] + rw [incl_of_counitAppApp] simp only [counitAppAppImage, CompHausLike.coe_of] letI : Fintype (fiber.{u, u} f x) := Fintype.ofInjective (sigmaIncl.{u, u} f x).1 Subtype.val_injective diff --git a/Mathlib/Geometry/RingedSpace/OpenImmersion.lean b/Mathlib/Geometry/RingedSpace/OpenImmersion.lean index 959214c257295..11dbae9b63d30 100644 --- a/Mathlib/Geometry/RingedSpace/OpenImmersion.lean +++ b/Mathlib/Geometry/RingedSpace/OpenImmersion.lean @@ -1120,8 +1120,8 @@ theorem lift_range (H' : Set.range g.base ⊆ Set.range f.base) : have : _ = (pullback.fst f g).base := PreservesPullback.iso_hom_fst (LocallyRingedSpace.forgetToSheafedSpace ⋙ SheafedSpace.forget _) f g - rw [LocallyRingedSpace.comp_base, ← this, ← Category.assoc, coe_comp] - rw [Set.range_comp, Set.range_iff_surjective.mpr, Set.image_univ] + rw [LocallyRingedSpace.comp_base, ← this, ← Category.assoc, coe_comp, Set.range_comp, + Set.range_iff_surjective.mpr, Set.image_univ] -- Porting note (#11224): change `rw` to `erw` on this lemma · erw [TopCat.pullback_fst_range] ext diff --git a/Mathlib/Topology/Order/ScottTopology.lean b/Mathlib/Topology/Order/ScottTopology.lean index 0a203e27ad789..2a58f0caeeb47 100644 --- a/Mathlib/Topology/Order/ScottTopology.lean +++ b/Mathlib/Topology/Order/ScottTopology.lean @@ -177,7 +177,7 @@ variable {α D} lemma isOpen_iff [IsScottHausdorff α D] : IsOpen s ↔ ∀ ⦃d : Set α⦄, d ∈ D → d.Nonempty → DirectedOn (· ≤ ·) d → ∀ ⦃a : α⦄, IsLUB d a → - a ∈ s → ∃ b ∈ d, Ici b ∩ d ⊆ s := by erw [topology_eq_scottHausdorff (α := α) (D := D)]; rfl + a ∈ s → ∃ b ∈ d, Ici b ∩ d ⊆ s := by rw [topology_eq_scottHausdorff (α := α) (D := D)]; rfl lemma dirSupInaccOn_of_isOpen [IsScottHausdorff α D] (h : IsOpen s) : DirSupInaccOn D s := fun d hd₀ hd₁ hd₂ a hda hd₃ ↦ by @@ -241,7 +241,7 @@ lemma topology_eq [IsScott α D] : ‹_› = scott α D := topology_eq_scott variable {α} {D} {s : Set α} {a : α} lemma isOpen_iff_isUpperSet_and_scottHausdorff_open [IsScott α D] : - IsOpen s ↔ IsUpperSet s ∧ IsOpen[scottHausdorff α D] s := by erw [topology_eq α D]; rfl + IsOpen s ↔ IsUpperSet s ∧ IsOpen[scottHausdorff α D] s := by rw [topology_eq α D]; rfl lemma isOpen_iff_isUpperSet_and_dirSupInaccOn [IsScott α D] : IsOpen s ↔ IsUpperSet s ∧ DirSupInaccOn D s := by From 7b3afdd0e7374470ac2af311f49a06197ede6db1 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 18 Oct 2024 16:28:16 +0000 Subject: [PATCH 152/505] feat: more lemmas about the size of pointwise sets (#17231) Complete the API and follow the convention that `Cardinal.mk` is called `card` and `Nat.card` is called `natCard` in lemma names. From LeanCamCombi --- Mathlib.lean | 1 + Mathlib/Algebra/Group/Pointwise/Set/Card.lean | 82 ++++++++++++------- .../GroupWithZero/Pointwise/Set/Card.lean | 27 ++++++ Mathlib/SetTheory/Cardinal/Basic.lean | 8 ++ 4 files changed, 89 insertions(+), 29 deletions(-) create mode 100644 Mathlib/Algebra/GroupWithZero/Pointwise/Set/Card.lean diff --git a/Mathlib.lean b/Mathlib.lean index 3c455bcceb7e9..b0b4fbb4d3724 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -330,6 +330,7 @@ import Mathlib.Algebra.GroupWithZero.NonZeroDivisors import Mathlib.Algebra.GroupWithZero.Opposite import Mathlib.Algebra.GroupWithZero.Pi import Mathlib.Algebra.GroupWithZero.Pointwise.Set.Basic +import Mathlib.Algebra.GroupWithZero.Pointwise.Set.Card import Mathlib.Algebra.GroupWithZero.Prod import Mathlib.Algebra.GroupWithZero.Semiconj import Mathlib.Algebra.GroupWithZero.ULift diff --git a/Mathlib/Algebra/Group/Pointwise/Set/Card.lean b/Mathlib/Algebra/Group/Pointwise/Set/Card.lean index 341ed441df932..80a073b792385 100644 --- a/Mathlib/Algebra/Group/Pointwise/Set/Card.lean +++ b/Mathlib/Algebra/Group/Pointwise/Set/Card.lean @@ -3,60 +3,84 @@ Copyright (c) 2024 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ -import Mathlib.Algebra.Group.Pointwise.Finset.Basic +import Mathlib.Data.Set.Pointwise.Finite import Mathlib.SetTheory.Cardinal.Finite /-! -# Cardinalities of pointwise operations on sets. +# Cardinalities of pointwise operations on sets -/ -namespace Set - -open Pointwise - -variable {α β : Type*} +open scoped Cardinal Pointwise -section MulAction -variable [Group α] [MulAction α β] +namespace Set +variable {G M α : Type*} -@[to_additive (attr := simp)] -lemma card_smul_set (a : α) (s : Set β) : Nat.card ↥(a • s) = Nat.card s := - Nat.card_image_of_injective (MulAction.injective a) _ +section Mul +variable [Mul M] {s t : Set M} -end MulAction +@[to_additive] +lemma _root_.Cardinal.mk_mul_le : #(s * t) ≤ #s * #t := by + rw [← image2_mul]; exact Cardinal.mk_image2_le -section IsCancelMul -variable [Mul α] [IsCancelMul α] {s t : Set α} +variable [IsCancelMul M] @[to_additive] -lemma card_mul_le : Nat.card (s * t) ≤ Nat.card s * Nat.card t := by - classical +lemma natCard_mul_le : Nat.card (s * t) ≤ Nat.card s * Nat.card t := by obtain h | h := (s * t).infinite_or_finite · simp [Set.Infinite.card_eq_zero h] - obtain ⟨hs, ht⟩ | rfl | rfl := finite_mul.1 h - · lift s to Finset α using hs - lift t to Finset α using ht - rw [← Finset.coe_mul] - simpa [-Finset.coe_mul] using Finset.card_mul_le - all_goals simp + simp only [Nat.card, ← Cardinal.toNat_mul] + refine Cardinal.toNat_le_toNat Cardinal.mk_mul_le ?_ + aesop (add simp [Cardinal.mul_lt_aleph0_iff, finite_mul]) + +@[to_additive (attr := deprecated (since := "2024-09-30"))] alias card_mul_le := natCard_mul_le -end IsCancelMul +end Mul section InvolutiveInv -variable [InvolutiveInv α] {s t : Set α} +variable [InvolutiveInv G] {s t : Set G} + +@[to_additive (attr := simp)] +lemma _root_.Cardinal.mk_inv (s : Set G) : #↥(s⁻¹) = #s := by + rw [← image_inv, Cardinal.mk_image_eq_of_injOn _ _ inv_injective.injOn] @[to_additive (attr := simp)] -lemma card_inv (s : Set α) : Nat.card ↥(s⁻¹) = Nat.card s := by +lemma natCard_inv (s : Set G) : Nat.card ↥(s⁻¹) = Nat.card s := by rw [← image_inv, Nat.card_image_of_injective inv_injective] +@[to_additive (attr := deprecated (since := "2024-09-30"))] alias card_inv := natCard_inv + end InvolutiveInv +section DivInvMonoid +variable [DivInvMonoid M] {s t : Set M} + +@[to_additive] +lemma _root_.Cardinal.mk_div_le : #(s / t) ≤ #s * #t := by + rw [← image2_div]; exact Cardinal.mk_image2_le + +end DivInvMonoid + section Group -variable [Group α] {s t : Set α} +variable [Group G] {s t : Set G} @[to_additive] -lemma card_div_le : Nat.card (s / t) ≤ Nat.card s * Nat.card t := by - rw [div_eq_mul_inv, ← card_inv t]; exact card_mul_le +lemma natCard_div_le : Nat.card (s / t) ≤ Nat.card s * Nat.card t := by + rw [div_eq_mul_inv, ← natCard_inv t]; exact natCard_mul_le + +@[to_additive (attr := deprecated (since := "2024-09-30"))] alias card_div_le := natCard_div_le + +variable [MulAction G α] + +@[to_additive (attr := simp)] +lemma _root_.Cardinal.mk_smul_set (a : G) (s : Set α) : #↥(a • s) = #s := + Cardinal.mk_image_eq_of_injOn _ _ (MulAction.injective a).injOn + +@[to_additive (attr := simp)] +lemma natCard_smul_set (a : G) (s : Set α) : Nat.card ↥(a • s) = Nat.card s := + Nat.card_image_of_injective (MulAction.injective a) _ + +@[to_additive (attr := deprecated (since := "2024-09-30"))] +alias card_smul_set := Cardinal.mk_smul_set end Group end Set diff --git a/Mathlib/Algebra/GroupWithZero/Pointwise/Set/Card.lean b/Mathlib/Algebra/GroupWithZero/Pointwise/Set/Card.lean new file mode 100644 index 0000000000000..8466c8876f1c1 --- /dev/null +++ b/Mathlib/Algebra/GroupWithZero/Pointwise/Set/Card.lean @@ -0,0 +1,27 @@ +/- +Copyright (c) 2024 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies +-/ +import Mathlib.Algebra.Group.Pointwise.Set.Basic +import Mathlib.Algebra.GroupWithZero.Action.Basic +import Mathlib.SetTheory.Cardinal.Finite + +/-! +# Cardinality of sets under pointwise group with zero operations +-/ + +open scoped Cardinal Pointwise + +variable {G G₀ M M₀ : Type*} + +namespace Set +variable [GroupWithZero G₀] [Zero M₀] [MulActionWithZero G₀ M₀] {a : G₀} + +lemma _root_.Cardinal.mk_smul_set₀ (ha : a ≠ 0) (s : Set M₀) : #↥(a • s) = #s := + Cardinal.mk_image_eq_of_injOn _ _ (MulAction.injective₀ ha).injOn + +lemma natCard_smul_set₀ (ha : a ≠ 0) (s : Set M₀) : Nat.card ↥(a • s) = Nat.card s := + Nat.card_image_of_injective (MulAction.injective₀ ha) _ + +end Set diff --git a/Mathlib/SetTheory/Cardinal/Basic.lean b/Mathlib/SetTheory/Cardinal/Basic.lean index 22ec5d425d8c1..fcc9e9174c5c2 100644 --- a/Mathlib/SetTheory/Cardinal/Basic.lean +++ b/Mathlib/SetTheory/Cardinal/Basic.lean @@ -1820,9 +1820,17 @@ theorem mk_emptyCollection_iff {α : Type u} {s : Set α} : #s = 0 ↔ s = ∅ : theorem mk_univ {α : Type u} : #(@univ α) = #α := mk_congr (Equiv.Set.univ α) +@[simp] lemma mk_setProd {α β : Type u} (s : Set α) (t : Set β) : #(s ×ˢ t) = #s * #t := by + rw [mul_def, mk_congr (Equiv.Set.prod ..)] + theorem mk_image_le {α β : Type u} {f : α → β} {s : Set α} : #(f '' s) ≤ #s := mk_le_of_surjective surjective_onto_image +lemma mk_image2_le {α β γ : Type u} {f : α → β → γ} {s : Set α} {t : Set β} : + #(image2 f s t) ≤ #s * #t := by + rw [← image_uncurry_prod, ← mk_setProd] + exact mk_image_le + theorem mk_image_le_lift {α : Type u} {β : Type v} {f : α → β} {s : Set α} : lift.{u} #(f '' s) ≤ lift.{v} #s := lift_mk_le.{0}.mpr ⟨Embedding.ofSurjective _ surjective_onto_image⟩ From 7e96f0c587bdf06977946b7227123bfd4ae81dc4 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Fri, 18 Oct 2024 16:28:18 +0000 Subject: [PATCH 153/505] chore: remove unused open statements (#17899) --- Mathlib/Algebra/Group/EvenFunction.lean | 2 -- Mathlib/Algebra/Polynomial/Module/Basic.lean | 2 +- Mathlib/Analysis/Calculus/LogDeriv.lean | 2 +- Mathlib/Analysis/Normed/Group/Ultra.lean | 2 -- .../Analysis/SpecialFunctions/Trigonometric/Cotangent.lean | 4 ++-- Mathlib/Combinatorics/Additive/ErdosGinzburgZiv.lean | 1 - Mathlib/Combinatorics/SetFamily/KruskalKatona.lean | 2 +- Mathlib/Data/Nat/Choose/Lucas.lean | 2 +- Mathlib/GroupTheory/CosetCover.lean | 2 +- Mathlib/GroupTheory/GroupAction/Blocks.lean | 2 +- Mathlib/MeasureTheory/Measure/SeparableMeasure.lean | 2 +- Mathlib/ModelTheory/Algebra/Field/IsAlgClosed.lean | 2 +- Mathlib/NumberTheory/Fermat.lean | 1 - Mathlib/NumberTheory/JacobiSum/Basic.lean | 2 +- Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean | 2 +- Mathlib/NumberTheory/NumberField/Units/Regulator.lean | 2 +- Mathlib/Order/LiminfLimsup.lean | 2 +- Mathlib/Probability/Martingale/BorelCantelli.lean | 2 +- Mathlib/RingTheory/RootsOfUnity/Lemmas.lean | 2 +- 19 files changed, 16 insertions(+), 22 deletions(-) diff --git a/Mathlib/Algebra/Group/EvenFunction.lean b/Mathlib/Algebra/Group/EvenFunction.lean index 3896969ce6795..4713a0549c61f 100644 --- a/Mathlib/Algebra/Group/EvenFunction.lean +++ b/Mathlib/Algebra/Group/EvenFunction.lean @@ -18,8 +18,6 @@ conflicting with the root-level definitions `Even` and `Odd` (which, for functio function takes even resp. odd _values_, a wholly different concept). -/ -open scoped BigOperators - namespace Function variable {α β : Type*} [Neg α] diff --git a/Mathlib/Algebra/Polynomial/Module/Basic.lean b/Mathlib/Algebra/Polynomial/Module/Basic.lean index 9a18a3951bc2e..c22c4aea59f64 100644 --- a/Mathlib/Algebra/Polynomial/Module/Basic.lean +++ b/Mathlib/Algebra/Polynomial/Module/Basic.lean @@ -14,7 +14,7 @@ This is defined as a type alias `PolynomialModule R M := ℕ →₀ M`, since th module structures on `ℕ →₀ M` of interest. See the docstring of `PolynomialModule` for details. -/ universe u v -open Polynomial BigOperators +open Polynomial /-- The `R[X]`-module `M[X]` for an `R`-module `M`. This is isomorphic (as an `R`-module) to `M[X]` when `M` is a ring. diff --git a/Mathlib/Analysis/Calculus/LogDeriv.lean b/Mathlib/Analysis/Calculus/LogDeriv.lean index 495f3de5c2775..7b8a8970a06f3 100644 --- a/Mathlib/Analysis/Calculus/LogDeriv.lean +++ b/Mathlib/Analysis/Calculus/LogDeriv.lean @@ -17,7 +17,7 @@ noncomputable section open Filter Function -open scoped Topology BigOperators Classical +open scoped Topology Classical variable {𝕜 𝕜': Type*} [NontriviallyNormedField 𝕜] [NontriviallyNormedField 𝕜'] [NormedAlgebra 𝕜 𝕜'] diff --git a/Mathlib/Analysis/Normed/Group/Ultra.lean b/Mathlib/Analysis/Normed/Group/Ultra.lean index 2b8c5f341691f..8881f035276fb 100644 --- a/Mathlib/Analysis/Normed/Group/Ultra.lean +++ b/Mathlib/Analysis/Normed/Group/Ultra.lean @@ -33,8 +33,6 @@ ultrametric, nonarchimedean -/ open Metric NNReal -open scoped BigOperators - namespace IsUltrametricDist section Group diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Cotangent.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Cotangent.lean index d9d5e56fb97ef..1db8bfb1259ff 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Cotangent.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Cotangent.lean @@ -11,9 +11,9 @@ import Mathlib.Analysis.Complex.UpperHalfPlane.Exp This file contains lemmas about the cotangent function, including useful series expansions. -/ -open Real Complex BigOperators Filter +open Real Complex -open scoped UpperHalfPlane Topology +open scoped UpperHalfPlane lemma Complex.cot_eq_exp_ratio (z : ℂ) : cot z = (Complex.exp (2 * I * z) + 1) / (I * (1 - Complex.exp (2 * I * z))) := by diff --git a/Mathlib/Combinatorics/Additive/ErdosGinzburgZiv.lean b/Mathlib/Combinatorics/Additive/ErdosGinzburgZiv.lean index a3bad041f207e..164ae582fd759 100644 --- a/Mathlib/Combinatorics/Additive/ErdosGinzburgZiv.lean +++ b/Mathlib/Combinatorics/Additive/ErdosGinzburgZiv.lean @@ -22,7 +22,6 @@ elements of sum zero. -/ open Finset MvPolynomial -open scoped BigOperators variable {ι : Type*} diff --git a/Mathlib/Combinatorics/SetFamily/KruskalKatona.lean b/Mathlib/Combinatorics/SetFamily/KruskalKatona.lean index 8b5ce0135fec7..23daec75ae3d2 100644 --- a/Mathlib/Combinatorics/SetFamily/KruskalKatona.lean +++ b/Mathlib/Combinatorics/SetFamily/KruskalKatona.lean @@ -119,7 +119,7 @@ protected lemma IsInitSeg.shadow [Finite α] (h₁ : IsInitSeg 𝒜 r) : IsInitS end Colex open Finset Colex Nat UV -open scoped BigOperators FinsetFamily +open scoped FinsetFamily variable {α : Type*} [LinearOrder α] {s U V : Finset α} {n : ℕ} diff --git a/Mathlib/Data/Nat/Choose/Lucas.lean b/Mathlib/Data/Nat/Choose/Lucas.lean index 94e97c1f452af..0df7d593456f8 100644 --- a/Mathlib/Data/Nat/Choose/Lucas.lean +++ b/Mathlib/Data/Nat/Choose/Lucas.lean @@ -22,7 +22,7 @@ k_i` modulo `p`, where `n_i` and `k_i` are the base-`p` digits of `n` and `k`, r open Finset hiding choose -open Nat BigOperators Polynomial +open Nat Polynomial namespace Choose diff --git a/Mathlib/GroupTheory/CosetCover.lean b/Mathlib/GroupTheory/CosetCover.lean index a8ba591868ffa..155e9fd467601 100644 --- a/Mathlib/GroupTheory/CosetCover.lean +++ b/Mathlib/GroupTheory/CosetCover.lean @@ -42,7 +42,7 @@ set of all minimal polynomials (not proved here). -/ -open scoped Pointwise BigOperators +open scoped Pointwise namespace Subgroup diff --git a/Mathlib/GroupTheory/GroupAction/Blocks.lean b/Mathlib/GroupTheory/GroupAction/Blocks.lean index e88c5972e23aa..cc63c90b25d16 100644 --- a/Mathlib/GroupTheory/GroupAction/Blocks.lean +++ b/Mathlib/GroupTheory/GroupAction/Blocks.lean @@ -43,7 +43,7 @@ We follow [wielandt1964]. -/ -open scoped BigOperators Pointwise +open scoped Pointwise namespace MulAction diff --git a/Mathlib/MeasureTheory/Measure/SeparableMeasure.lean b/Mathlib/MeasureTheory/Measure/SeparableMeasure.lean index 167c0d5bdfb53..cb0e5e2a0c352 100644 --- a/Mathlib/MeasureTheory/Measure/SeparableMeasure.lean +++ b/Mathlib/MeasureTheory/Measure/SeparableMeasure.lean @@ -63,7 +63,7 @@ written `≠ ∞` rather than `< ∞`. See `Ne.lt_top` and `ne_of_lt` to switch separable measure, measure-dense, Lp space, second-countable -/ -open MeasurableSpace Set ENNReal TopologicalSpace BigOperators symmDiff Filter Real +open MeasurableSpace Set ENNReal TopologicalSpace symmDiff Real namespace MeasureTheory diff --git a/Mathlib/ModelTheory/Algebra/Field/IsAlgClosed.lean b/Mathlib/ModelTheory/Algebra/Field/IsAlgClosed.lean index 2d7efbda8ea35..54708abe18b61 100644 --- a/Mathlib/ModelTheory/Algebra/Field/IsAlgClosed.lean +++ b/Mathlib/ModelTheory/Algebra/Field/IsAlgClosed.lean @@ -49,7 +49,7 @@ namespace FirstOrder namespace Field -open Ring FreeCommRing BigOperators Polynomial Language +open Ring FreeCommRing Polynomial Language /-- A generic monic polynomial of degree `n` as an element of the free commutative ring in `n+1` variables, with a variable for each diff --git a/Mathlib/NumberTheory/Fermat.lean b/Mathlib/NumberTheory/Fermat.lean index c90f62582cae8..0e3ab436b2414 100644 --- a/Mathlib/NumberTheory/Fermat.lean +++ b/Mathlib/NumberTheory/Fermat.lean @@ -26,7 +26,6 @@ for all natural numbers `n`. namespace Nat open Finset -open scoped BigOperators /-- Fermat numbers: the `n`-th Fermat number is defined as `2^(2^n) + 1`. -/ def fermatNumber (n : ℕ) : ℕ := 2 ^ (2 ^ n) + 1 diff --git a/Mathlib/NumberTheory/JacobiSum/Basic.lean b/Mathlib/NumberTheory/JacobiSum/Basic.lean index a936f799462ca..7e6fd3e597e49 100644 --- a/Mathlib/NumberTheory/JacobiSum/Basic.lean +++ b/Mathlib/NumberTheory/JacobiSum/Basic.lean @@ -28,7 +28,7 @@ but generalize where appropriate. This is based on Lean code written as part of the bachelor's thesis of Alexander Spahl. -/ -open BigOperators Finset +open Finset /-! ### Jacobi sums: definition and first properties diff --git a/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean b/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean index 4970ba0001a5a..cc34b948b40e0 100644 --- a/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean +++ b/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean @@ -40,7 +40,7 @@ open scoped NumberField noncomputable section -open NumberField NumberField.InfinitePlace NumberField.Units BigOperators +open NumberField NumberField.InfinitePlace NumberField.Units variable (K : Type*) [Field K] diff --git a/Mathlib/NumberTheory/NumberField/Units/Regulator.lean b/Mathlib/NumberTheory/NumberField/Units/Regulator.lean index 34d7ce2339491..cd7cb76b5bf07 100644 --- a/Mathlib/NumberTheory/NumberField/Units/Regulator.lean +++ b/Mathlib/NumberTheory/NumberField/Units/Regulator.lean @@ -32,7 +32,7 @@ namespace NumberField.Units variable (K : Type*) [Field K] -open MeasureTheory Classical BigOperators NumberField.InfinitePlace +open MeasureTheory Classical NumberField.InfinitePlace NumberField NumberField.Units.dirichletUnitTheorem variable [NumberField K] diff --git a/Mathlib/Order/LiminfLimsup.lean b/Mathlib/Order/LiminfLimsup.lean index 0d6faea4d46f6..300e1d0160937 100644 --- a/Mathlib/Order/LiminfLimsup.lean +++ b/Mathlib/Order/LiminfLimsup.lean @@ -285,7 +285,7 @@ end Relation section add_and_sum -open Filter BigOperators Set +open Filter Set variable {α : Type*} {f : Filter α} variable {R : Type*} diff --git a/Mathlib/Probability/Martingale/BorelCantelli.lean b/Mathlib/Probability/Martingale/BorelCantelli.lean index 4897ac3d82a8e..0f483a8596a0b 100644 --- a/Mathlib/Probability/Martingale/BorelCantelli.lean +++ b/Mathlib/Probability/Martingale/BorelCantelli.lean @@ -34,7 +34,7 @@ and `ProbabilityTheory.measure_limsup_eq_one` for the second (which does). open Filter -open scoped NNReal ENNReal MeasureTheory ProbabilityTheory BigOperators Topology +open scoped NNReal ENNReal MeasureTheory ProbabilityTheory Topology namespace MeasureTheory diff --git a/Mathlib/RingTheory/RootsOfUnity/Lemmas.lean b/Mathlib/RingTheory/RootsOfUnity/Lemmas.lean index e6fded4815f54..33d3afe29785a 100644 --- a/Mathlib/RingTheory/RootsOfUnity/Lemmas.lean +++ b/Mathlib/RingTheory/RootsOfUnity/Lemmas.lean @@ -23,7 +23,7 @@ variable {R : Type*} [CommRing R] [IsDomain R] namespace IsPrimitiveRoot -open Finset Polynomial BigOperators +open Finset Polynomial /-- If `μ` is a primitive `n`th root of unity in `R`, then `∏(1≤k Date: Fri, 18 Oct 2024 16:28:19 +0000 Subject: [PATCH 154/505] chore: remove `CoeFun` instances where `FunLike` is available (#17900) During the port we found that `FunLike` is robust enough not to need an extra `CoeFun` shortcut. Let's see if that rule can be consistently applied to the whole of the library. There is still duplication between `FunLike` and `CoeFun` for `Grp`, `Mon`, `CommGrp` and `CommMon`, which will need a more thorough fix. See also #17866. This is the second half of #17870, and doesn't affect the benchmarks. --- Mathlib/Algebra/Algebra/Hom.lean | 2 -- Mathlib/Algebra/Algebra/NonUnitalHom.lean | 5 ----- Mathlib/Algebra/MonoidAlgebra/Defs.lean | 4 ++-- Mathlib/Algebra/Polynomial/Module/Basic.lean | 2 +- Mathlib/LinearAlgebra/AffineSpace/AffineEquiv.lean | 3 --- Mathlib/LinearAlgebra/AffineSpace/AffineMap.lean | 5 ----- Mathlib/LinearAlgebra/QuadraticForm/Basic.lean | 5 ----- Mathlib/MeasureTheory/OuterMeasure/Defs.lean | 3 --- Mathlib/ModelTheory/Basic.lean | 6 ------ Mathlib/ModelTheory/ElementaryMaps.lean | 3 --- Mathlib/Order/OmegaCompletePartialOrder.lean | 1 - Mathlib/Topology/Algebra/Module/Multilinear/Basic.lean | 3 --- Mathlib/Topology/Algebra/Module/WeakDual.lean | 5 ----- Mathlib/Topology/Connected/PathConnected.lean | 6 ------ Mathlib/Topology/Homeomorph.lean | 2 -- Mathlib/Topology/MetricSpace/Dilation.lean | 3 --- Mathlib/Topology/MetricSpace/DilationEquiv.lean | 3 --- 17 files changed, 3 insertions(+), 58 deletions(-) diff --git a/Mathlib/Algebra/Algebra/Hom.lean b/Mathlib/Algebra/Algebra/Hom.lean index 55dcdc81a91cb..5c5eac38616b5 100644 --- a/Mathlib/Algebra/Algebra/Hom.lean +++ b/Mathlib/Algebra/Algebra/Hom.lean @@ -85,8 +85,6 @@ section Semiring variable [CommSemiring R] [Semiring A] [Semiring B] [Semiring C] [Semiring D] variable [Algebra R A] [Algebra R B] [Algebra R C] [Algebra R D] --- Porting note: we don't port specialized `CoeFun` instances if there is `DFunLike` instead - instance funLike : FunLike (A →ₐ[R] B) A B where coe f := f.toFun coe_injective' f g h := by diff --git a/Mathlib/Algebra/Algebra/NonUnitalHom.lean b/Mathlib/Algebra/Algebra/NonUnitalHom.lean index 34a5518dc4de3..7eb27b49ee74a 100644 --- a/Mathlib/Algebra/Algebra/NonUnitalHom.lean +++ b/Mathlib/Algebra/Algebra/NonUnitalHom.lean @@ -156,11 +156,6 @@ variable [NonUnitalNonAssocSemiring A] [DistribMulAction R A] variable [NonUnitalNonAssocSemiring B] [DistribMulAction S B] variable [NonUnitalNonAssocSemiring C] [DistribMulAction T C] --- Porting note: Replaced with DFunLike instance --- /-- see Note [function coercion] -/ --- instance : CoeFun (A →ₙₐ[R] B) fun _ => A → B := --- ⟨toFun⟩ - instance : DFunLike (A →ₛₙₐ[φ] B) A fun _ => B where coe f := f.toFun coe_injective' := by rintro ⟨⟨⟨f, _⟩, _⟩, _⟩ ⟨⟨⟨g, _⟩, _⟩, _⟩ h; congr diff --git a/Mathlib/Algebra/MonoidAlgebra/Defs.lean b/Mathlib/Algebra/MonoidAlgebra/Defs.lean index 7125b7d636d97..767b5198c52d8 100644 --- a/Mathlib/Algebra/MonoidAlgebra/Defs.lean +++ b/Mathlib/Algebra/MonoidAlgebra/Defs.lean @@ -86,7 +86,7 @@ instance MonoidAlgebra.instIsCancelAdd [IsCancelAdd k] : IsCancelAdd (MonoidAlge inferInstanceAs (IsCancelAdd (G →₀ k)) instance MonoidAlgebra.coeFun : CoeFun (MonoidAlgebra k G) fun _ => G → k := - Finsupp.instCoeFun + inferInstanceAs (CoeFun (G →₀ k) _) end @@ -823,7 +823,7 @@ instance instIsCancelAdd [IsCancelAdd k] : IsCancelAdd (AddMonoidAlgebra k G) := inferInstanceAs (IsCancelAdd (G →₀ k)) instance coeFun : CoeFun k[G] fun _ => G → k := - Finsupp.instCoeFun + inferInstanceAs (CoeFun (G →₀ k) _) end AddMonoidAlgebra diff --git a/Mathlib/Algebra/Polynomial/Module/Basic.lean b/Mathlib/Algebra/Polynomial/Module/Basic.lean index c22c4aea59f64..047d5c7f86864 100644 --- a/Mathlib/Algebra/Polynomial/Module/Basic.lean +++ b/Mathlib/Algebra/Polynomial/Module/Basic.lean @@ -57,7 +57,7 @@ instance instFunLike : FunLike (PolynomialModule R M) ℕ M := Finsupp.instFunLike instance : CoeFun (PolynomialModule R M) fun _ => ℕ → M := - Finsupp.instCoeFun + inferInstanceAs <| CoeFun (_ →₀ _) _ theorem zero_apply (i : ℕ) : (0 : PolynomialModule R M) i = 0 := Finsupp.zero_apply diff --git a/Mathlib/LinearAlgebra/AffineSpace/AffineEquiv.lean b/Mathlib/LinearAlgebra/AffineSpace/AffineEquiv.lean index 0849ae203d5a8..3dcf47e22249d 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/AffineEquiv.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/AffineEquiv.lean @@ -89,9 +89,6 @@ instance equivLike : EquivLike (P₁ ≃ᵃ[k] P₂) P₁ P₂ where right_inv f := f.right_inv coe_injective' _ _ h _ := toAffineMap_injective (DFunLike.coe_injective h) -instance : CoeFun (P₁ ≃ᵃ[k] P₂) fun _ => P₁ → P₂ := - DFunLike.hasCoeToFun - instance : CoeOut (P₁ ≃ᵃ[k] P₂) (P₁ ≃ P₂) := ⟨AffineEquiv.toEquiv⟩ diff --git a/Mathlib/LinearAlgebra/AffineSpace/AffineMap.lean b/Mathlib/LinearAlgebra/AffineSpace/AffineMap.lean index edeab9e15af30..71f99e5c85dde 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/AffineMap.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/AffineMap.lean @@ -67,11 +67,6 @@ instance AffineMap.instFunLike (k : Type*) {V1 : Type*} (P1 : Type*) {V2 : Type* apply vadd_right_cancel (f p) rw [← f_add, h, ← g_add] -instance AffineMap.hasCoeToFun (k : Type*) {V1 : Type*} (P1 : Type*) {V2 : Type*} (P2 : Type*) - [Ring k] [AddCommGroup V1] [Module k V1] [AffineSpace V1 P1] [AddCommGroup V2] [Module k V2] - [AffineSpace V2 P2] : CoeFun (P1 →ᵃ[k] P2) fun _ => P1 → P2 := - DFunLike.hasCoeToFun - namespace LinearMap variable {k : Type*} {V₁ : Type*} {V₂ : Type*} [Ring k] [AddCommGroup V₁] [Module k V₁] diff --git a/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean b/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean index fe2fe2b66dfe2..39bc41327d511 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean @@ -161,11 +161,6 @@ instance instFunLike : FunLike (QuadraticMap R M N) M N where coe := toFun coe_injective' x y h := by cases x; cases y; congr -/-- Helper instance for when there's too many metavariables to apply -`DFunLike.hasCoeToFun` directly. -/ -instance : CoeFun (QuadraticMap R M N) fun _ => M → N := - ⟨DFunLike.coe⟩ - variable (Q) /-- The `simp` normal form for a quadratic map is `DFunLike.coe`, not `toFun`. -/ diff --git a/Mathlib/MeasureTheory/OuterMeasure/Defs.lean b/Mathlib/MeasureTheory/OuterMeasure/Defs.lean index e1169e732155e..a776ff90e2c9c 100644 --- a/Mathlib/MeasureTheory/OuterMeasure/Defs.lean +++ b/Mathlib/MeasureTheory/OuterMeasure/Defs.lean @@ -62,9 +62,6 @@ instance : FunLike (OuterMeasure α) (Set α) ℝ≥0∞ where coe m := m.measureOf coe_injective' | ⟨_, _, _, _⟩, ⟨_, _, _, _⟩, rfl => rfl -instance instCoeFun : CoeFun (OuterMeasure α) (fun _ => Set α → ℝ≥0∞) := - inferInstance - @[simp] theorem measureOf_eq_coe (m : OuterMeasure α) : m.measureOf = m := rfl instance : OuterMeasureClass (OuterMeasure α) α where diff --git a/Mathlib/ModelTheory/Basic.lean b/Mathlib/ModelTheory/Basic.lean index 6846aee6eb65c..08419dda7c6f0 100644 --- a/Mathlib/ModelTheory/Basic.lean +++ b/Mathlib/ModelTheory/Basic.lean @@ -296,9 +296,6 @@ instance homClass : HomClass L (M →[L] N) M N where instance [L.IsAlgebraic] : StrongHomClass L (M →[L] N) M N := HomClass.strongHomClassOfIsAlgebraic -instance hasCoeToFun : CoeFun (M →[L] N) fun _ => M → N := - DFunLike.hasCoeToFun - @[simp] theorem toFun_eq_coe {f : M →[L] N} : f.toFun = (f : M → N) := rfl @@ -557,9 +554,6 @@ def symm (f : M ≃[L] N) : N ≃[L] M := refine (f.map_rel' r (f.toEquiv.symm ∘ x)).symm.trans ?_ rw [← Function.comp_assoc, Equiv.toFun_as_coe, Equiv.self_comp_symm, Function.id_comp] } -instance hasCoeToFun : CoeFun (M ≃[L] N) fun _ => M → N := - DFunLike.hasCoeToFun - @[simp] theorem symm_symm (f : M ≃[L] N) : f.symm.symm = f := diff --git a/Mathlib/ModelTheory/ElementaryMaps.lean b/Mathlib/ModelTheory/ElementaryMaps.lean index f774228f18d50..a8c9fc2716718 100644 --- a/Mathlib/ModelTheory/ElementaryMaps.lean +++ b/Mathlib/ModelTheory/ElementaryMaps.lean @@ -67,9 +67,6 @@ instance instFunLike : FunLike (M ↪ₑ[L] N) M N where ext x exact funext_iff.1 h x -instance : CoeFun (M ↪ₑ[L] N) fun _ => M → N := - DFunLike.hasCoeToFun - @[simp] theorem map_boundedFormula (f : M ↪ₑ[L] N) {α : Type*} {n : ℕ} (φ : L.BoundedFormula α n) (v : α → M) (xs : Fin n → M) : φ.Realize (f ∘ v) (f ∘ xs) ↔ φ.Realize v xs := by diff --git a/Mathlib/Order/OmegaCompletePartialOrder.lean b/Mathlib/Order/OmegaCompletePartialOrder.lean index 0d7b0cf6a3804..179530a5796f4 100644 --- a/Mathlib/Order/OmegaCompletePartialOrder.lean +++ b/Mathlib/Order/OmegaCompletePartialOrder.lean @@ -72,7 +72,6 @@ variable [Preorder α] [Preorder β] [Preorder γ] instance : FunLike (Chain α) ℕ α := inferInstanceAs <| FunLike (ℕ →o α) ℕ α instance : OrderHomClass (Chain α) ℕ α := inferInstanceAs <| OrderHomClass (ℕ →o α) ℕ α -instance : CoeFun (Chain α) fun _ => ℕ → α := ⟨DFunLike.coe⟩ instance [Inhabited α] : Inhabited (Chain α) := ⟨⟨default, fun _ _ _ => le_rfl⟩⟩ diff --git a/Mathlib/Topology/Algebra/Module/Multilinear/Basic.lean b/Mathlib/Topology/Algebra/Module/Multilinear/Basic.lean index ad413d2ba21ae..1a0a5179e3380 100644 --- a/Mathlib/Topology/Algebra/Module/Multilinear/Basic.lean +++ b/Mathlib/Topology/Algebra/Module/Multilinear/Basic.lean @@ -78,9 +78,6 @@ instance continuousMapClass : ContinuousMapClass (ContinuousMultilinearMap R M₁ M₂) (∀ i, M₁ i) M₂ where map_continuous := ContinuousMultilinearMap.cont -instance : CoeFun (ContinuousMultilinearMap R M₁ M₂) fun _ => (∀ i, M₁ i) → M₂ := - ⟨fun f => f⟩ - /-- See Note [custom simps projection]. We need to specify this projection explicitly in this case, because it is a composition of multiple projections. -/ def Simps.apply (L₁ : ContinuousMultilinearMap R M₁ M₂) (v : ∀ i, M₁ i) : M₂ := diff --git a/Mathlib/Topology/Algebra/Module/WeakDual.lean b/Mathlib/Topology/Algebra/Module/WeakDual.lean index 1a71efa833d83..45e6b09628962 100644 --- a/Mathlib/Topology/Algebra/Module/WeakDual.lean +++ b/Mathlib/Topology/Algebra/Module/WeakDual.lean @@ -102,11 +102,6 @@ instance instFunLike : FunLike (WeakDual 𝕜 E) E 𝕜 := instance instContinuousLinearMapClass : ContinuousLinearMapClass (WeakDual 𝕜 E) 𝕜 E 𝕜 := ContinuousLinearMap.continuousSemilinearMapClass -/-- Helper instance for when there's too many metavariables to apply `DFunLike.hasCoeToFun` -directly. -/ -instance : CoeFun (WeakDual 𝕜 E) fun _ => E → 𝕜 := - DFunLike.hasCoeToFun - /-- If a monoid `M` distributively continuously acts on `𝕜` and this action commutes with multiplication on `𝕜`, then it acts on `WeakDual 𝕜 E`. -/ instance instMulAction (M) [Monoid M] [DistribMulAction M 𝕜] [SMulCommClass 𝕜 M 𝕜] diff --git a/Mathlib/Topology/Connected/PathConnected.lean b/Mathlib/Topology/Connected/PathConnected.lean index 4cc61d6513e73..93ef11a0f5e30 100644 --- a/Mathlib/Topology/Connected/PathConnected.lean +++ b/Mathlib/Topology/Connected/PathConnected.lean @@ -85,12 +85,6 @@ instance Path.funLike : FunLike (Path x y) I X where instance Path.continuousMapClass : ContinuousMapClass (Path x y) I X where map_continuous γ := show Continuous γ.toContinuousMap by fun_prop --- Porting note: not necessary in light of the instance above -/- -instance : CoeFun (Path x y) fun _ => I → X := - ⟨fun p => p.toFun⟩ --/ - @[ext] protected theorem Path.ext : ∀ {γ₁ γ₂ : Path x y}, (γ₁ : I → X) = γ₂ → γ₁ = γ₂ := by rintro ⟨⟨x, h11⟩, h12, h13⟩ ⟨⟨x, h21⟩, h22, h23⟩ rfl diff --git a/Mathlib/Topology/Homeomorph.lean b/Mathlib/Topology/Homeomorph.lean index 09be530f7a37a..c406c1e326c58 100644 --- a/Mathlib/Topology/Homeomorph.lean +++ b/Mathlib/Topology/Homeomorph.lean @@ -60,8 +60,6 @@ instance : EquivLike (X ≃ₜ Y) X Y where right_inv h := h.right_inv coe_injective' _ _ H _ := toEquiv_injective <| DFunLike.ext' H -instance : CoeFun (X ≃ₜ Y) fun _ ↦ X → Y := ⟨DFunLike.coe⟩ - @[simp] theorem homeomorph_mk_coe (a : X ≃ Y) (b c) : (Homeomorph.mk a b c : X → Y) = a := rfl diff --git a/Mathlib/Topology/MetricSpace/Dilation.lean b/Mathlib/Topology/MetricSpace/Dilation.lean index 2622d97a69454..0ece9c35abd15 100644 --- a/Mathlib/Topology/MetricSpace/Dilation.lean +++ b/Mathlib/Topology/MetricSpace/Dilation.lean @@ -87,9 +87,6 @@ instance funLike : FunLike (α →ᵈ β) α β where instance toDilationClass : DilationClass (α →ᵈ β) α β where edist_eq' f := edist_eq' f -instance : CoeFun (α →ᵈ β) fun _ => α → β := - DFunLike.hasCoeToFun - @[simp] theorem toFun_eq_coe {f : α →ᵈ β} : f.toFun = (f : α → β) := rfl diff --git a/Mathlib/Topology/MetricSpace/DilationEquiv.lean b/Mathlib/Topology/MetricSpace/DilationEquiv.lean index d510cd58fb9e7..3fffba7e08305 100644 --- a/Mathlib/Topology/MetricSpace/DilationEquiv.lean +++ b/Mathlib/Topology/MetricSpace/DilationEquiv.lean @@ -60,9 +60,6 @@ instance : EquivLike (X ≃ᵈ Y) X Y where instance : DilationEquivClass (X ≃ᵈ Y) X Y where edist_eq' f := f.edist_eq' -instance : CoeFun (X ≃ᵈ Y) fun _ ↦ (X → Y) where - coe f := f - @[simp] theorem coe_toEquiv (e : X ≃ᵈ Y) : ⇑e.toEquiv = e := rfl @[ext] From 2178cceeb3452f1a474caa729442569a4646056d Mon Sep 17 00:00:00 2001 From: mathlib4-update-dependencies-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Fri, 18 Oct 2024 16:57:13 +0000 Subject: [PATCH 155/505] chore: update Mathlib dependencies 2024-10-18 (#17922) This PR updates the Mathlib dependencies. --- lake-manifest.json | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lake-manifest.json b/lake-manifest.json index b86cb2dfc481e..c93ba44db832e 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -15,7 +15,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "fa3d73a2cf077f4b14c7840352ac7b08aeb6eb41", + "rev": "1357f4f49450abb9dfd4783e38219f4ce84f9785", "name": "Qq", "manifestFile": "lake-manifest.json", "inputRev": "master", From 1ef8dd30084da1bb42462baaf2dbfcafc07701c9 Mon Sep 17 00:00:00 2001 From: Scott Carnahan <128885296+ScottCarnahan@users.noreply.github.com> Date: Fri, 18 Oct 2024 17:35:48 +0000 Subject: [PATCH 156/505] feat (RingTheory/HahnSeries/Summable) : singleton summable family (#16871) This PR defines a summable family made of a single Hahn series, and shows that the formal sum of the family is equal to the original Hahn series. --- Mathlib/RingTheory/HahnSeries/Summable.lean | 13 +++++++++++++ 1 file changed, 13 insertions(+) diff --git a/Mathlib/RingTheory/HahnSeries/Summable.lean b/Mathlib/RingTheory/HahnSeries/Summable.lean index 45a21761c48c3..a4a57ee6c1624 100644 --- a/Mathlib/RingTheory/HahnSeries/Summable.lean +++ b/Mathlib/RingTheory/HahnSeries/Summable.lean @@ -178,6 +178,19 @@ theorem hsum_add {s t : SummableFamily Γ R α} : (s + t).hsum = s.hsum + t.hsum simp only [hsum_coeff, add_coeff, add_apply] exact finsum_add_distrib (s.finite_co_support _) (t.finite_co_support _) +/-- The summable family made of a single Hahn series. -/ +@[simps] +def single (x : HahnSeries Γ R) : SummableFamily Γ R Unit where + toFun _ := x + isPWO_iUnion_support' := + Eq.mpr (congrArg (fun s ↦ s.IsPWO) (Set.iUnion_const x.support)) x.isPWO_support + finite_co_support' g := Set.toFinite {a | ((fun _ ↦ x) a).coeff g ≠ 0} + +@[simp] +theorem hsum_single (x : HahnSeries Γ R) : (single x).hsum = x := by + ext g + simp only [hsum_coeff, single_toFun, finsum_unique] + /-- A summable family induced by an equivalence of the parametrizing type. -/ @[simps] def Equiv (e : α ≃ β) (s : SummableFamily Γ R α) : SummableFamily Γ R β where From 75204348c7d435ac1d22d07c32f5d2836d9548ed Mon Sep 17 00:00:00 2001 From: Thmoas-Guan <150537269+Thmoas-Guan@users.noreply.github.com> Date: Fri, 18 Oct 2024 17:54:26 +0000 Subject: [PATCH 157/505] feat(FieldTheory/Galois): Finite galois subextensions (#16978) Define the type of finite Galois subextensions and prove basic properties about it. Co-authored-by: Yuyang Zhao Nailin Guan <3571819596@qq.com> Jujian Zhang Moves: - Mathlib.FieldTheory.Galois -> Mathlib.FieldTheory.Galois.Basic --- Mathlib.lean | 1 + Mathlib/FieldTheory/Galois/Infinite.lean | 214 +++++++++++++++++++++++ Mathlib/FieldTheory/KrullTopology.lean | 11 ++ Mathlib/FieldTheory/Normal.lean | 45 +++++ Mathlib/FieldTheory/NormalClosure.lean | 17 ++ 5 files changed, 288 insertions(+) create mode 100644 Mathlib/FieldTheory/Galois/Infinite.lean diff --git a/Mathlib.lean b/Mathlib.lean index b0b4fbb4d3724..c66a3419a9aeb 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2724,6 +2724,7 @@ import Mathlib.FieldTheory.Finite.Trace import Mathlib.FieldTheory.Finiteness import Mathlib.FieldTheory.Fixed import Mathlib.FieldTheory.Galois.Basic +import Mathlib.FieldTheory.Galois.Infinite import Mathlib.FieldTheory.IntermediateField.Algebraic import Mathlib.FieldTheory.IntermediateField.Basic import Mathlib.FieldTheory.IsAlgClosed.AlgebraicClosure diff --git a/Mathlib/FieldTheory/Galois/Infinite.lean b/Mathlib/FieldTheory/Galois/Infinite.lean new file mode 100644 index 0000000000000..d3c2345887e0d --- /dev/null +++ b/Mathlib/FieldTheory/Galois/Infinite.lean @@ -0,0 +1,214 @@ +/- +Copyright (c) 2024 Nailin Guan. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Nailin Guan, Yuyang Zhao, Jujian Zhang +-/ + +import Mathlib.Algebra.Category.Grp.FiniteGrp +import Mathlib.CategoryTheory.Category.Preorder +import Mathlib.FieldTheory.NormalClosure +import Mathlib.FieldTheory.SeparableClosure + +/-! + +# Main definitions and results + +In a field extension `K/k` + +* `FiniteGaloisIntermediateField` : The type of intermediate fields of `K/k` + that are finite and Galois over `k` + +* `adjoin` : The finite Galois intermediate field obtained from the normal closure of adjoining a + finite `s : Set K` to `k`. + +* `finGaloisGroup L` : The (finite) Galois group `Gal(L/k)` associated to a + `L : FiniteGaloisIntermediateField k K` `L`. + +* `finGaloisGroupMap` : For `FiniteGaloisIntermediateField` s `L₁` and `L₂` with `L₂ ≤ L₁` + giving the restriction of `Gal(L₁/k)` to `Gal(L₂/k)` + +* `finGaloisGroupFunctor` : Mapping `FiniteGaloisIntermediateField` ordered by inverse inclusion to + its corresponding Galois Group as `FiniteGrp`. + +# TODO + +* `FiniteGaloisIntermediateField` should be a `ConditionallyCompleteLattice` but isn't proved yet. + +-/ + +open CategoryTheory Opposite +open scoped IntermediateField + +variable (k K : Type*) [Field k] [Field K] [Algebra k K] + +/-- The type of intermediate fields of `K/k` that are finite and Galois over `k` -/ +structure FiniteGaloisIntermediateField extends IntermediateField k K where + [finiteDimensional : FiniteDimensional k toIntermediateField] + [isGalois : IsGalois k toIntermediateField] + +namespace FiniteGaloisIntermediateField + +instance : Coe (FiniteGaloisIntermediateField k K) (IntermediateField k K) where + coe := toIntermediateField + +instance : CoeSort (FiniteGaloisIntermediateField k K) (Type _) where + coe L := L.toIntermediateField + +instance (L : FiniteGaloisIntermediateField k K) : FiniteDimensional k L := + L.finiteDimensional + +instance (L : FiniteGaloisIntermediateField k K) : IsGalois k L := L.isGalois + +variable {k K} + +lemma val_injective : Function.Injective (toIntermediateField (k := k) (K := K)) := by + rintro ⟨⟩ ⟨⟩ eq + simpa only [mk.injEq] using eq + +/-- Turns the collection of finite Galois IntermediateFields of `K/k` into a lattice. -/ + +instance (L₁ L₂ : IntermediateField k K) [IsGalois k L₁] [IsGalois k L₂] : + IsGalois k ↑(L₁ ⊔ L₂) where + +instance (L₁ L₂ : IntermediateField k K) [FiniteDimensional k L₁] : + FiniteDimensional k ↑(L₁ ⊓ L₂) := + .of_injective (IntermediateField.inclusion inf_le_left).toLinearMap + (IntermediateField.inclusion inf_le_left).injective + +instance (L₁ L₂ : IntermediateField k K) [FiniteDimensional k L₂] : + FiniteDimensional k ↑(L₁ ⊓ L₂) := + .of_injective (IntermediateField.inclusion inf_le_right).toLinearMap + (IntermediateField.inclusion inf_le_right).injective + +instance (L₁ L₂ : IntermediateField k K) [Algebra.IsSeparable k L₁] : + Algebra.IsSeparable k ↑(L₁ ⊓ L₂) := + .of_algHom _ _ (IntermediateField.inclusion inf_le_left) + +instance (L₁ L₂ : IntermediateField k K) [Algebra.IsSeparable k L₂] : + Algebra.IsSeparable k ↑(L₁ ⊓ L₂) := + .of_algHom _ _ (IntermediateField.inclusion inf_le_right) + +instance (L₁ L₂ : IntermediateField k K) [IsGalois k L₁] [IsGalois k L₂] : + IsGalois k ↑(L₁ ⊓ L₂) where + +instance : Sup (FiniteGaloisIntermediateField k K) where + sup L₁ L₂ := .mk <| L₁ ⊔ L₂ + +instance : Inf (FiniteGaloisIntermediateField k K) where + inf L₁ L₂ := .mk <| L₁ ⊓ L₂ + +instance : Lattice (FiniteGaloisIntermediateField k K) := + val_injective.lattice _ (fun _ _ ↦ rfl) (fun _ _ ↦ rfl) + +instance : OrderBot (FiniteGaloisIntermediateField k K) where + bot := .mk ⊥ + bot_le _ := bot_le (α := IntermediateField _ _) + +@[simp] +lemma le_iff (L₁ L₂ : FiniteGaloisIntermediateField k K) : + L₁ ≤ L₂ ↔ L₁.toIntermediateField ≤ L₂.toIntermediateField := + Iff.rfl + +variable (k) in +/-- The minimal (finite) Galois intermediate field containing a finite set `s : Set K` in a +Galois extension `K/k` defined as the the normal closure of the field obtained by adjoining +the set `s : Set K` to `k`. -/ +noncomputable def adjoin [IsGalois k K] (s : Set K) [Finite s] : + FiniteGaloisIntermediateField k K := { + normalClosure k (IntermediateField.adjoin k (s : Set K)) K with + finiteDimensional := + letI : FiniteDimensional k (IntermediateField.adjoin k (s : Set K)) := + IntermediateField.finiteDimensional_adjoin <| fun z _ => + IsAlgebraic.isIntegral (Algebra.IsAlgebraic.isAlgebraic z) + normalClosure.is_finiteDimensional k (IntermediateField.adjoin k (s : Set K)) K + isGalois := IsGalois.normalClosure k (IntermediateField.adjoin k (s : Set K)) K } + +@[simp] +lemma adjoin_val [IsGalois k K] (s : Set K) [Finite s] : + (FiniteGaloisIntermediateField.adjoin k s) = + normalClosure k (IntermediateField.adjoin k s) K := + rfl + +variable (k) in +lemma subset_adjoin [IsGalois k K] (s : Set K) [Finite s] : + s ⊆ (adjoin k s).toIntermediateField := + (IntermediateField.subset_adjoin k s).trans (IntermediateField.le_normalClosure _) + +theorem adjoin_simple_le_iff [IsGalois k K] {x : K} {L : FiniteGaloisIntermediateField k K} : + adjoin k {x} ≤ L ↔ x ∈ L.toIntermediateField := by + simp only [le_iff, adjoin_val, IntermediateField.normalClosure_le_iff_of_normal, + IntermediateField.adjoin_le_iff, Set.le_eq_subset, Set.singleton_subset_iff, SetLike.mem_coe] + +@[simp] +theorem adjoin_map [IsGalois k K] (f : K →ₐ[k] K) (s : Set K) [Finite s] : + adjoin k (f '' s) = adjoin k s := by + apply val_injective; dsimp [adjoin_val] + rw [← IntermediateField.adjoin_map, IntermediateField.normalClosure_map_eq] + +@[simp] +theorem adjoin_simple_map_algHom [IsGalois k K] (f : K →ₐ[k] K) (x : K) : + adjoin k {f x} = adjoin k {x} := by + simpa only [Set.image_singleton] using adjoin_map f { x } + +@[simp] +theorem adjoin_simple_map_algEquiv [IsGalois k K] (f : K ≃ₐ[k] K) (x : K) : + adjoin k {f x} = adjoin k {x} := + adjoin_simple_map_algHom (f : K →ₐ[k] K) x + +end FiniteGaloisIntermediateField + +section Profinite + +variable {k K} + +/-- The (finite) Galois group `Gal(L / k)` associated to a +`L : FiniteGaloisIntermediateField k K` `L`. -/ +def FiniteGaloisIntermediateField.finGaloisGroup (L : FiniteGaloisIntermediateField k K) : + FiniteGrp := + letI := AlgEquiv.fintype k L + FiniteGrp.of <| L ≃ₐ[k] L + +/-- For `FiniteGaloisIntermediateField` s `L₁` and `L₂` with `L₂ ≤ L₁` + the restriction homomorphism from `Gal(L₁/k)` to `Gal(L₂/k)` -/ +noncomputable def finGaloisGroupMap {L₁ L₂ : (FiniteGaloisIntermediateField k K)ᵒᵖ} + (le : L₁ ⟶ L₂) : L₁.unop.finGaloisGroup ⟶ L₂.unop.finGaloisGroup := + haveI : Normal k L₂.unop := IsGalois.to_normal + letI : Algebra L₂.unop L₁.unop := RingHom.toAlgebra (Subsemiring.inclusion <| leOfHom le.1) + haveI : IsScalarTower k L₂.unop L₁.unop := IsScalarTower.of_algebraMap_eq (congrFun rfl) + FiniteGrp.ofHom (AlgEquiv.restrictNormalHom L₂.unop) + +namespace finGaloisGroupMap + +@[simp] +lemma map_id (L : (FiniteGaloisIntermediateField k K)ᵒᵖ) : + (finGaloisGroupMap (𝟙 L)) = 𝟙 L.unop.finGaloisGroup := + AlgEquiv.restrictNormalHom_id _ _ + +@[simp] +lemma map_comp {L₁ L₂ L₃ : (FiniteGaloisIntermediateField k K)ᵒᵖ} (f : L₁ ⟶ L₂) (g : L₂ ⟶ L₃) : + finGaloisGroupMap (f ≫ g) = finGaloisGroupMap f ≫ finGaloisGroupMap g := by + iterate 2 + induction L₁ with | _ L₁ => ?_ + induction L₂ with | _ L₂ => ?_ + induction L₃ with | _ L₃ => ?_ + letI : Algebra L₃ L₂ := RingHom.toAlgebra (Subsemiring.inclusion g.unop.le) + letI : Algebra L₂ L₁ := RingHom.toAlgebra (Subsemiring.inclusion f.unop.le) + letI : Algebra L₃ L₁ := RingHom.toAlgebra (Subsemiring.inclusion (g.unop.le.trans f.unop.le)) + haveI : IsScalarTower k L₂ L₁ := IsScalarTower.of_algebraMap_eq' rfl + haveI : IsScalarTower k L₃ L₁ := IsScalarTower.of_algebraMap_eq' rfl + haveI : IsScalarTower k L₃ L₂ := IsScalarTower.of_algebraMap_eq' rfl + haveI : IsScalarTower L₃ L₂ L₁ := IsScalarTower.of_algebraMap_eq' rfl + apply IsScalarTower.AlgEquiv.restrictNormalHom_comp k L₃ L₂ L₁ + +end finGaloisGroupMap + +variable (k K) in +/-- The functor from `FiniteGaloisIntermediateField` (ordered by reverse inclusion) to `FiniteGrp`, +mapping each intermediate field `K/L/k` to `Gal (L/k)`.-/ +noncomputable def finGaloisGroupFunctor : (FiniteGaloisIntermediateField k K)ᵒᵖ ⥤ FiniteGrp where + obj L := L.unop.finGaloisGroup + map := finGaloisGroupMap + map_id := finGaloisGroupMap.map_id + map_comp := finGaloisGroupMap.map_comp + +end Profinite diff --git a/Mathlib/FieldTheory/KrullTopology.lean b/Mathlib/FieldTheory/KrullTopology.lean index 7f7f53c3109a7..c5751d36c7ff4 100644 --- a/Mathlib/FieldTheory/KrullTopology.lean +++ b/Mathlib/FieldTheory/KrullTopology.lean @@ -178,6 +178,17 @@ instance krullTopology (K L : Type*) [Field K] [Field L] [Algebra K L] : instance (K L : Type*) [Field K] [Field L] [Algebra K L] : TopologicalGroup (L ≃ₐ[K] L) := GroupFilterBasis.isTopologicalGroup (galGroupBasis K L) +open scoped Topology in +lemma krullTopology_mem_nhds_one (K L : Type*) [Field K] [Field L] [Algebra K L] + (s : Set (L ≃ₐ[K] L)) : s ∈ 𝓝 1 ↔ ∃ E : IntermediateField K L, + FiniteDimensional K E ∧ (E.fixingSubgroup : Set (L ≃ₐ[K] L)) ⊆ s := by + rw [GroupFilterBasis.nhds_one_eq] + constructor + · rintro ⟨-, ⟨-, ⟨E, fin, rfl⟩, rfl⟩, hE⟩ + exact ⟨E, fin, hE⟩ + · rintro ⟨E, fin, hE⟩ + exact ⟨E.fixingSubgroup, ⟨E.fixingSubgroup, ⟨E, fin, rfl⟩, rfl⟩, hE⟩ + section KrullT2 open scoped Topology Filter diff --git a/Mathlib/FieldTheory/Normal.lean b/Mathlib/FieldTheory/Normal.lean index 98ae051a5fd57..242f0ef840a7d 100644 --- a/Mathlib/FieldTheory/Normal.lean +++ b/Mathlib/FieldTheory/Normal.lean @@ -355,10 +355,55 @@ theorem AlgEquiv.restrict_liftNormal (χ : K₁ ≃ₐ[F] K₁) [Normal F K₁] (algebraMap K₁ E).injective (Eq.trans (AlgEquiv.restrictNormal_commutes _ K₁ x) (χ.liftNormal_commutes E x)) +/-- The group homomorphism given by restricting an algebra isomorphism to a normal subfield +is surjective. -/ theorem AlgEquiv.restrictNormalHom_surjective [Normal F K₁] [Normal F E] : Function.Surjective (AlgEquiv.restrictNormalHom K₁ : (E ≃ₐ[F] E) → K₁ ≃ₐ[F] K₁) := fun χ => ⟨χ.liftNormal E, χ.restrict_liftNormal E⟩ +/-- The group homomorphism given by restricting an algebra isomorphism to itself +is the identity map. -/ +@[simp] +theorem AlgEquiv.restrictNormalHom_id (F K : Type*) + [Field F] [Field K] [Algebra F K] [Normal F K] : + AlgEquiv.restrictNormalHom K = MonoidHom.id (K ≃ₐ[F] K) := by + ext f x + dsimp only [restrictNormalHom, MonoidHom.mk'_apply, MonoidHom.id_apply] + apply (algebraMap K K).injective + rw [AlgEquiv.restrictNormal_commutes] + simp only [Algebra.id.map_eq_id, RingHom.id_apply] + +namespace IsScalarTower + +/-- In a scalar tower `K₃/K₂/K₁/F` with `K₁` and `K₂` are normal over `F`, the group homomorphism +given by the restriction of algebra isomorphisms of `K₃` to `K₁` is equal to the composition of +the group homomorphism given by the restricting an algebra isomorphism of `K₃` to `K₂` and +the group homomorphism given by the restricting an algebra isomorphism of `K₂` to `K₁` -/ +theorem AlgEquiv.restrictNormalHom_comp (F K₁ K₂ K₃ : Type*) + [Field F] [Field K₁] [Field K₂] [Field K₃] + [Algebra F K₁] [Algebra F K₂] [Algebra F K₃] [Algebra K₁ K₂] [Algebra K₁ K₃] [Algebra K₂ K₃] + [IsScalarTower F K₁ K₃] [IsScalarTower F K₁ K₂] [IsScalarTower F K₂ K₃] [IsScalarTower K₁ K₂ K₃] + [Normal F K₁] [Normal F K₂] : + AlgEquiv.restrictNormalHom K₁ = + (AlgEquiv.restrictNormalHom K₁).comp + (AlgEquiv.restrictNormalHom (F := F) (K₁ := K₃) K₂) := by + ext f x + apply (algebraMap K₁ K₃).injective + rw [IsScalarTower.algebraMap_eq K₁ K₂ K₃] + simp only [AlgEquiv.restrictNormalHom, MonoidHom.mk'_apply, RingHom.coe_comp, Function.comp_apply, + ← algebraMap_apply, AlgEquiv.restrictNormal_commutes, MonoidHom.coe_comp] + +theorem AlgEquiv.restrictNormalHom_comp_apply (K₁ K₂ : Type*) {F K₃ : Type*} + [Field F] [Field K₁] [Field K₂] [Field K₃] + [Algebra F K₁] [Algebra F K₂] [Algebra F K₃] [Algebra K₁ K₂] [Algebra K₁ K₃] [Algebra K₂ K₃] + [IsScalarTower F K₁ K₃] [IsScalarTower F K₁ K₂] [IsScalarTower F K₂ K₃] [IsScalarTower K₁ K₂ K₃] + [Normal F K₁] [Normal F K₂] (f : K₃ ≃ₐ[F] K₃) : + AlgEquiv.restrictNormalHom K₁ f = + (AlgEquiv.restrictNormalHom K₁) (AlgEquiv.restrictNormalHom K₂ f) := by + rw [IsScalarTower.AlgEquiv.restrictNormalHom_comp F K₁ K₂ K₃, MonoidHom.comp_apply] + +end IsScalarTower + open IntermediateField in theorem Normal.minpoly_eq_iff_mem_orbit [h : Normal F E] {x y : E} : minpoly F x = minpoly F y ↔ x ∈ MulAction.orbit (E ≃ₐ[F] E) y := by diff --git a/Mathlib/FieldTheory/NormalClosure.lean b/Mathlib/FieldTheory/NormalClosure.lean index 680dc94bf8709..fac28822a66a5 100644 --- a/Mathlib/FieldTheory/NormalClosure.lean +++ b/Mathlib/FieldTheory/NormalClosure.lean @@ -282,4 +282,21 @@ lemma normal_iff_forall_map_eq : Normal F K ↔ ∀ σ : L →ₐ[F] L, K.map σ lemma normal_iff_forall_map_eq' : Normal F K ↔ ∀ σ : L ≃ₐ[F] L, K.map ↑σ = K := ⟨fun h σ ↦ normal_iff_forall_map_eq.1 h σ, fun h ↦ normal_iff_forall_map_le'.2 (fun σ ↦ (h σ).le)⟩ +@[simp] +lemma normalClosure_map_eq (K : IntermediateField F L) (σ : L →ₐ[F] L) : + normalClosure F (K.map σ) L = normalClosure F K L := by + have (σ : L ≃ₐ[F] L) : normalClosure F (K.map (σ : L →ₐ[F] L)) L = normalClosure F K L := by + simp_rw [normalClosure_def'', map_map] + exact (Equiv.mulRight σ).iSup_congr fun _ ↦ rfl + exact this ((Algebra.IsAlgebraic.algEquivEquivAlgHom _ _).symm σ) + +@[simp] +theorem normalClosure_le_iff_of_normal {K₁ K₂ : IntermediateField F L} [Normal F K₂] : + normalClosure F K₁ L ≤ K₂ ↔ K₁ ≤ K₂ := by + refine ⟨fun h ↦ ?_, fun h ↦ ?_⟩ + · rw [normalClosure_le_iff] at h + simpa only [fieldRange_val] using h K₁.val + · rw [← normalClosure_of_normal K₂] + exact normalClosure_mono K₁ K₂ h + end IntermediateField From 8610319f04e6b637e3c597905fa5cb405e5549fe Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Fri, 18 Oct 2024 18:06:40 +0000 Subject: [PATCH 158/505] =?UTF-8?q?doc(SetTheory/Game/PGame):=20expand=20d?= =?UTF-8?q?ocstrings=20for=20`=E2=89=A4`=20and=20`=E2=A7=8F`=20(#16809)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- Mathlib/SetTheory/Game/PGame.lean | 11 ++++++++--- 1 file changed, 8 insertions(+), 3 deletions(-) diff --git a/Mathlib/SetTheory/Game/PGame.lean b/Mathlib/SetTheory/Game/PGame.lean index afa7e0b220917..83f9d1f8d2559 100644 --- a/Mathlib/SetTheory/Game/PGame.lean +++ b/Mathlib/SetTheory/Game/PGame.lean @@ -23,6 +23,9 @@ takes two types (thought of as indexing the possible moves for the players Left pair of functions out of these types to `SetTheory.PGame` (thought of as describing the resulting game after making a move). +We may denote a game as $\{L | R\}$, where $L$ and $R$ stand for the collections of left and right +moves. This notation is not currently used in Mathlib. + Combinatorial games themselves, as a quotient of pregames, are constructed in `Game.lean`. ## Conway induction @@ -349,15 +352,17 @@ instance isEmpty_one_rightMoves : IsEmpty (RightMoves 1) := /-- The less or equal relation on pre-games. -If `0 ≤ x`, then Left can win `x` as the second player. -/ +If `0 ≤ x`, then Left can win `x` as the second player. `x ≤ y` means that `0 ≤ y - x`. +See `PGame.le_iff_sub_nonneg`. -/ instance le : LE PGame := ⟨Sym2.GameAdd.fix wf_isOption fun x y le => (∀ i, ¬le y (x.moveLeft i) (Sym2.GameAdd.snd_fst <| IsOption.moveLeft i)) ∧ ∀ j, ¬le (y.moveRight j) x (Sym2.GameAdd.fst_snd <| IsOption.moveRight j)⟩ -/-- The less or fuzzy relation on pre-games. +/-- The less or fuzzy relation on pre-games. `x ⧏ y` is defined as `¬ y ≤ x`. -If `0 ⧏ x`, then Left can win `x` as the first player. -/ +If `0 ⧏ x`, then Left can win `x` as the first player. `x ⧏ y` means that `0 ⧏ y - x`. +See `PGame.lf_iff_sub_zero_lf`. -/ def LF (x y : PGame) : Prop := ¬y ≤ x From df601e35b5257a5c391ee5ca5648201c4f46cc78 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Fri, 18 Oct 2024 18:31:29 +0000 Subject: [PATCH 159/505] feat(AlgebraicGeometry/EllipticCurve/Projective): implement group operations on point representatives for projective coordinates (#9418) Define the analogous secant-and-tangent negation `neg` and addition `add` on `PointRep` over `F`, and lift them to `PointClass`. Define `Point` as a `PointClass` that is nonsingular. Prove in `neg_equiv` and `add_equiv` that these operations preserve the equivalence. Prove in `nonsingular_neg` and `nonsingular_add` that these operations preserve nonsingularity by reducing it to the affine case, and lift these proofs to `PointClass`. This is the third in a series of four PRs leading to #8485 --- .../EllipticCurve/Projective.lean | 282 +++++++++++++++++- scripts/no_lints_prime_decls.txt | 1 + 2 files changed, 282 insertions(+), 1 deletion(-) diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/Projective.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/Projective.lean index 3d4ea8337acc9..00a1a91f99235 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/Projective.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/Projective.lean @@ -14,7 +14,8 @@ import Mathlib.Tactic.LinearCombination' This file defines the type of points on a Weierstrass curve as a tuple, consisting of an equivalence class of triples up to scaling by a unit, satisfying a Weierstrass equation with a nonsingular -condition. +condition. This file also defines the negation and addition operations of the group law for this +type, and proves that they respect the Weierstrass equation and the nonsingular condition. ## Mathematical background @@ -32,6 +33,7 @@ given by a tuple consisting of $[x:y:z]$ and the nonsingular condition on any re As in `Mathlib.AlgebraicGeometry.EllipticCurve.Affine`, the set of nonsingular rational points forms an abelian group under the same secant-and-tangent process, but the polynomials involved are homogeneous, and any instances of division become multiplication in the $Z$-coordinate. +Note that most computational proofs follow from their analogous proofs for affine coordinates. ## Main definitions @@ -39,10 +41,16 @@ homogeneous, and any instances of division become multiplication in the $Z$-coor * `WeierstrassCurve.Projective.toAffine`: the Weierstrass curve in affine coordinates. * `WeierstrassCurve.Projective.Nonsingular`: the nonsingular condition on a point representative. * `WeierstrassCurve.Projective.NonsingularLift`: the nonsingular condition on a point class. + * `WeierstrassCurve.Projective.neg`: the negation operation on a point representative. + * `WeierstrassCurve.Projective.negMap`: the negation operation on a point class. + * `WeierstrassCurve.Projective.add`: the addition operation on a point representative. + * `WeierstrassCurve.Projective.addMap`: the addition operation on a point class. ## Main statements * `WeierstrassCurve.Projective.polynomial_relation`: Euler's homogeneous function theorem. + * `WeierstrassCurve.Projective.nonsingular_neg`: negation preserves the nonsingular condition. + * `WeierstrassCurve.Projective.nonsingular_add`: addition preserves the nonsingular condition. ## Implementation notes @@ -138,6 +146,11 @@ lemma smul_equiv (P : Fin 3 → R) {u : R} (hu : IsUnit u) : u • P ≈ P := lemma smul_eq (P : Fin 3 → R) {u : R} (hu : IsUnit u) : (⟦u • P⟧ : PointClass R) = ⟦P⟧ := Quotient.eq.mpr <| smul_equiv P hu +lemma smul_equiv_smul (P Q : Fin 3 → R) {u v : R} (hu : IsUnit u) (hv : IsUnit v) : + u • P ≈ v • Q ↔ P ≈ Q := by + erw [← Quotient.eq_iff_equiv, ← Quotient.eq_iff_equiv, smul_eq P hu, smul_eq Q hv] + rfl + variable (W') in /-- The coercion to a Weierstrass curve in affine coordinates. -/ abbrev toAffine : Affine R := @@ -1163,4 +1176,271 @@ lemma addXYZ_of_Z_ne_zero {P Q : Fin 3 → F} (hP : W.Equation P) (hQ : W.Equati end Addition +section Negation + +/-! ### Negation on point representatives -/ + +variable (W') in +/-- The negation of a point representative. -/ +def neg (P : Fin 3 → R) : Fin 3 → R := + ![P x, W'.negY P, P z] + +lemma neg_X (P : Fin 3 → R) : W'.neg P x = P x := + rfl + +lemma neg_Y (P : Fin 3 → R) : W'.neg P y = W'.negY P := + rfl + +lemma neg_Z (P : Fin 3 → R) : W'.neg P z = P z := + rfl + +lemma neg_smul (P : Fin 3 → R) (u : R) : W'.neg (u • P) = u • W'.neg P := by + simpa only [neg, negY_smul] using (smul_fin3 (W'.neg P) u).symm + +lemma neg_smul_equiv (P : Fin 3 → R) {u : R} (hu : IsUnit u) : W'.neg (u • P) ≈ W'.neg P := + ⟨hu.unit, (neg_smul ..).symm⟩ + +lemma neg_equiv {P Q : Fin 3 → R} (h : P ≈ Q) : W'.neg P ≈ W'.neg Q := by + rcases h with ⟨u, rfl⟩ + exact neg_smul_equiv Q u.isUnit + +lemma neg_of_Z_eq_zero [NoZeroDivisors R] {P : Fin 3 → R} (hP : W'.Equation P) (hPz : P z = 0) : + W'.neg P = -P y • ![0, 1, 0] := by + erw [neg, X_eq_zero_of_Z_eq_zero hP hPz, negY_of_Z_eq_zero hP hPz, hPz, smul_fin3, mul_zero, + mul_one] + +lemma neg_of_Z_ne_zero {P : Fin 3 → F} (hPz : P z ≠ 0) : + W.neg P = P z • ![P x / P z, W.toAffine.negY (P x / P z) (P y / P z), 1] := by + erw [neg, smul_fin3, mul_div_cancel₀ _ hPz, ← negY_of_Z_ne_zero hPz, mul_div_cancel₀ _ hPz, + mul_one] + +private lemma nonsingular_neg_of_Z_ne_zero {P : Fin 3 → F} (hP : W.Nonsingular P) (hPz : P z ≠ 0) : + W.Nonsingular ![P x / P z, W.toAffine.negY (P x / P z) (P y / P z), 1] := + (nonsingular_some ..).mpr <| Affine.nonsingular_neg <| (nonsingular_of_Z_ne_zero hPz).mp hP + +lemma nonsingular_neg {P : Fin 3 → F} (hP : W.Nonsingular P) : W.Nonsingular <| W.neg P := by + by_cases hPz : P z = 0 + · simp only [neg_of_Z_eq_zero hP.left hPz, nonsingular_smul _ (isUnit_Y_of_Z_eq_zero hP hPz).neg, + nonsingular_zero] + · simp only [neg_of_Z_ne_zero hPz, nonsingular_smul _ <| Ne.isUnit hPz, + nonsingular_neg_of_Z_ne_zero hP hPz] + +lemma addZ_neg (P : Fin 3 → R) : W'.addZ P (W'.neg P) = 0 := by + rw [addZ, neg_X, neg_Y, neg_Z, negY] + ring1 + +lemma addX_neg (P : Fin 3 → R) : W'.addX P (W'.neg P) = 0 := by + rw [addX, neg_X, neg_Y, neg_Z, negY] + ring1 + +lemma negAddY_neg {P : Fin 3 → R} (hP : W'.Equation P) : W'.negAddY P (W'.neg P) = W'.dblZ P := by + linear_combination (norm := (rw [negAddY, neg_X, neg_Y, neg_Z, dblZ, negY]; ring1)) + -3 * (P y - W'.negY P) * (equation_iff _).mp hP + +lemma addY_neg {P : Fin 3 → R} (hP : W'.Equation P) : W'.addY P (W'.neg P) = -W'.dblZ P := by + rw [addY, negY_eq, addX_neg, negAddY_neg hP, addZ_neg, mul_zero, sub_zero, mul_zero, sub_zero] + +lemma addXYZ_neg {P : Fin 3 → R} (hP : W'.Equation P) : + W'.addXYZ P (W'.neg P) = -W'.dblZ P • ![0, 1, 0] := by + erw [addXYZ, addX_neg, addY_neg hP, addZ_neg, smul_fin3, mul_zero, mul_one] + +variable (W') in +/-- The negation of a point class. If `P` is a point representative, +then `W'.negMap ⟦P⟧` is definitionally equivalent to `W'.neg P`. -/ +def negMap (P : PointClass R) : PointClass R := + P.map W'.neg fun _ _ => neg_equiv + +lemma negMap_eq (P : Fin 3 → R) : W'.negMap ⟦P⟧ = ⟦W'.neg P⟧ := + rfl + +lemma negMap_of_Z_eq_zero {P : Fin 3 → F} (hP : W.Nonsingular P) (hPz : P z = 0) : + W.negMap ⟦P⟧ = ⟦![0, 1, 0]⟧ := by + rw [negMap_eq, neg_of_Z_eq_zero hP.left hPz, smul_eq _ (isUnit_Y_of_Z_eq_zero hP hPz).neg] + +lemma negMap_of_Z_ne_zero {P : Fin 3 → F} (hPz : P z ≠ 0) : + W.negMap ⟦P⟧ = ⟦![P x / P z, W.toAffine.negY (P x / P z) (P y / P z), 1]⟧ := by + rw [negMap_eq, neg_of_Z_ne_zero hPz, smul_eq _ <| Ne.isUnit hPz] + +lemma nonsingularLift_negMap {P : PointClass F} (hP : W.NonsingularLift P) : + W.NonsingularLift <| W.negMap P := by + rcases P with ⟨_⟩ + exact nonsingular_neg hP + +end Negation + +section Addition + +/-! ### Addition on point representatives -/ + +open Classical in +variable (W') in +/-- The addition of two point representatives. -/ +noncomputable def add (P Q : Fin 3 → R) : Fin 3 → R := + if P ≈ Q then W'.dblXYZ P else W'.addXYZ P Q + +lemma add_of_equiv {P Q : Fin 3 → R} (h : P ≈ Q) : W'.add P Q = W'.dblXYZ P := + if_pos h + +lemma add_smul_of_equiv {P Q : Fin 3 → R} (h : P ≈ Q) {u v : R} (hu : IsUnit u) (hv : IsUnit v) : + W'.add (u • P) (v • Q) = u ^ 4 • W'.add P Q := by + rw [add_of_equiv <| (smul_equiv_smul P Q hu hv).mpr h, dblXYZ_smul, add_of_equiv h] + +lemma add_self (P : Fin 3 → R) : W'.add P P = W'.dblXYZ P := + add_of_equiv <| Setoid.refl _ + +lemma add_of_eq {P Q : Fin 3 → R} (h : P = Q) : W'.add P Q = W'.dblXYZ P := + h ▸ add_self P + +lemma add_of_not_equiv {P Q : Fin 3 → R} (h : ¬P ≈ Q) : W'.add P Q = W'.addXYZ P Q := + if_neg h + +lemma add_smul_of_not_equiv {P Q : Fin 3 → R} (h : ¬P ≈ Q) {u v : R} (hu : IsUnit u) + (hv : IsUnit v) : W'.add (u • P) (v • Q) = (u * v) ^ 2 • W'.add P Q := by + rw [add_of_not_equiv <| h.comp (smul_equiv_smul P Q hu hv).mp, addXYZ_smul, add_of_not_equiv h] + +lemma add_smul_equiv (P Q : Fin 3 → R) {u v : R} (hu : IsUnit u) (hv : IsUnit v) : + W'.add (u • P) (v • Q) ≈ W'.add P Q := by + by_cases h : P ≈ Q + · exact ⟨hu.unit ^ 4, by convert (add_smul_of_equiv h hu hv).symm⟩ + · exact ⟨(hu.unit * hv.unit) ^ 2, by convert (add_smul_of_not_equiv h hu hv).symm⟩ + +lemma add_equiv {P P' Q Q' : Fin 3 → R} (hP : P ≈ P') (hQ : Q ≈ Q') : + W'.add P Q ≈ W'.add P' Q' := by + rcases hP, hQ with ⟨⟨u, rfl⟩, ⟨v, rfl⟩⟩ + exact add_smul_equiv P' Q' u.isUnit v.isUnit + +lemma add_of_Z_eq_zero {P Q : Fin 3 → F} (hP : W.Nonsingular P) (hQ : W.Nonsingular Q) + (hPz : P z = 0) (hQz : Q z = 0) : W.add P Q = P y ^ 4 • ![0, 1, 0] := by + rw [add, if_pos <| equiv_of_Z_eq_zero hP hQ hPz hQz, dblXYZ_of_Z_eq_zero hP.left hPz] + +lemma add_of_Z_eq_zero_left [NoZeroDivisors R] {P Q : Fin 3 → R} (hP : W'.Equation P) + (hPz : P z = 0) (hQz : Q z ≠ 0) : W'.add P Q = (P y ^ 2 * Q z) • Q := by + rw [add, if_neg <| not_equiv_of_Z_eq_zero_left hPz hQz, addXYZ_of_Z_eq_zero_left hP hPz] + +lemma add_of_Z_eq_zero_right [NoZeroDivisors R] {P Q : Fin 3 → R} (hQ : W'.Equation Q) + (hPz : P z ≠ 0) (hQz : Q z = 0) : W'.add P Q = -(Q y ^ 2 * P z) • P := by + rw [add, if_neg <| not_equiv_of_Z_eq_zero_right hPz hQz, addXYZ_of_Z_eq_zero_right hQ hQz] + +lemma add_of_Y_eq {P Q : Fin 3 → F} (hP : W.Equation P) (hPz : P z ≠ 0) (hQz : Q z ≠ 0) + (hx : P x * Q z = Q x * P z) (hy : P y * Q z = Q y * P z) (hy' : P y * Q z = W.negY Q * P z) : + W.add P Q = W.dblU P • ![0, 1, 0] := by + rw [add, if_pos <| equiv_of_X_eq_of_Y_eq hPz hQz hx hy, dblXYZ_of_Y_eq hP hPz hQz hx hy hy'] + +lemma add_of_Y_ne {P Q : Fin 3 → F} (hP : W.Equation P) (hQ : W.Equation Q) (hPz : P z ≠ 0) + (hQz : Q z ≠ 0) (hx : P x * Q z = Q x * P z) (hy : P y * Q z ≠ Q y * P z) : + W.add P Q = addU P Q • ![0, 1, 0] := by + rw [add, if_neg <| not_equiv_of_Y_ne hy, addXYZ_of_X_eq hP hQ hPz hQz hx] + +lemma add_of_Y_ne' {P Q : Fin 3 → F} (hP : W.Equation P) (hQ : W.Equation Q) (hPz : P z ≠ 0) + (hQz : Q z ≠ 0) (hx : P x * Q z = Q x * P z) (hy : P y * Q z ≠ W.negY Q * P z) : + W.add P Q = W.dblZ P • + ![W.toAffine.addX (P x / P z) (Q x / Q z) + (W.toAffine.slope (P x / P z) (Q x / Q z) (P y / P z) (Q y / Q z)), + W.toAffine.addY (P x / P z) (Q x / Q z) (P y / P z) + (W.toAffine.slope (P x / P z) (Q x / Q z) (P y / P z) (Q y / Q z)), 1] := by + rw [add, if_pos <| equiv_of_X_eq_of_Y_eq hPz hQz hx <| Y_eq_of_Y_ne' hP hQ hPz hQz hx hy, + dblXYZ_of_Z_ne_zero hP hQ hPz hQz hx hy] + +lemma add_of_X_ne {P Q : Fin 3 → F} (hP : W.Equation P) (hQ : W.Equation Q) (hPz : P z ≠ 0) + (hQz : Q z ≠ 0) (hx : P x * Q z ≠ Q x * P z) : W.add P Q = W.addZ P Q • + ![W.toAffine.addX (P x / P z) (Q x / Q z) + (W.toAffine.slope (P x / P z) (Q x / Q z) (P y / P z) (Q y / Q z)), + W.toAffine.addY (P x / P z) (Q x / Q z) (P y / P z) + (W.toAffine.slope (P x / P z) (Q x / Q z) (P y / P z) (Q y / Q z)), 1] := by + rw [add, if_neg <| not_equiv_of_X_ne hx, addXYZ_of_Z_ne_zero hP hQ hPz hQz hx] + +private lemma nonsingular_add_of_Z_ne_zero {P Q : Fin 3 → F} (hP : W.Nonsingular P) + (hQ : W.Nonsingular Q) (hPz : P z ≠ 0) (hQz : Q z ≠ 0) + (hxy : P x * Q z = Q x * P z → P y * Q z ≠ W.negY Q * P z) : W.Nonsingular + ![W.toAffine.addX (P x / P z) (Q x / Q z) + (W.toAffine.slope (P x / P z) (Q x / Q z) (P y / P z) (Q y / Q z)), + W.toAffine.addY (P x / P z) (Q x / Q z) (P y / P z) + (W.toAffine.slope (P x / P z) (Q x / Q z) (P y / P z) (Q y / Q z)), 1] := + (nonsingular_some ..).mpr <| Affine.nonsingular_add ((nonsingular_of_Z_ne_zero hPz).mp hP) + ((nonsingular_of_Z_ne_zero hQz).mp hQ) (by rwa [← X_eq_iff hPz hQz, ne_eq, ← Y_eq_iff' hPz hQz]) + +lemma nonsingular_add {P Q : Fin 3 → F} (hP : W.Nonsingular P) (hQ : W.Nonsingular Q) : + W.Nonsingular <| W.add P Q := by + by_cases hPz : P z = 0 + · by_cases hQz : Q z = 0 + · simp only [add_of_Z_eq_zero hP hQ hPz hQz, + nonsingular_smul _ <| (isUnit_Y_of_Z_eq_zero hP hPz).pow 4, nonsingular_zero] + · simpa only [add_of_Z_eq_zero_left hP.left hPz hQz, + nonsingular_smul _ <| ((isUnit_Y_of_Z_eq_zero hP hPz).pow 2).mul <| Ne.isUnit hQz] + · by_cases hQz : Q z = 0 + · simpa only [add_of_Z_eq_zero_right hQ.left hPz hQz, + nonsingular_smul _ (((isUnit_Y_of_Z_eq_zero hQ hQz).pow 2).mul <| Ne.isUnit hPz).neg] + · by_cases hxy : P x * Q z = Q x * P z → P y * Q z ≠ W.negY Q * P z + · by_cases hx : P x * Q z = Q x * P z + · simp only [add_of_Y_ne' hP.left hQ.left hPz hQz hx <| hxy hx, + nonsingular_smul _ <| isUnit_dblZ_of_Y_ne' hP.left hQ.left hPz hQz hx <| hxy hx, + nonsingular_add_of_Z_ne_zero hP hQ hPz hQz hxy] + · simp only [add_of_X_ne hP.left hQ.left hPz hQz hx, + nonsingular_smul _ <| isUnit_addZ_of_X_ne hP.left hQ.left hx, + nonsingular_add_of_Z_ne_zero hP hQ hPz hQz hxy] + · rw [_root_.not_imp, not_ne_iff] at hxy + by_cases hy : P y * Q z = Q y * P z + · simp only [add_of_Y_eq hP.left hPz hQz hxy.left hy hxy.right, nonsingular_smul _ <| + isUnit_dblU_of_Y_eq hP hPz hQz hxy.left hy hxy.right, nonsingular_zero] + · simp only [add_of_Y_ne hP.left hQ.left hPz hQz hxy.left hy, + nonsingular_smul _ <| isUnit_addU_of_Y_ne hPz hQz hy, nonsingular_zero] + +variable (W') in +/-- The addition of two point classes. If `P` is a point representative, +then `W.addMap ⟦P⟧ ⟦Q⟧` is definitionally equivalent to `W.add P Q`. -/ +noncomputable def addMap (P Q : PointClass R) : PointClass R := + Quotient.map₂ W'.add (fun _ _ hP _ _ hQ => add_equiv hP hQ) P Q + +lemma addMap_eq (P Q : Fin 3 → R) : W'.addMap ⟦P⟧ ⟦Q⟧ = ⟦W'.add P Q⟧ := + rfl + +lemma addMap_of_Z_eq_zero_left {P : Fin 3 → F} {Q : PointClass F} (hP : W.Nonsingular P) + (hQ : W.NonsingularLift Q) (hPz : P z = 0) : W.addMap ⟦P⟧ Q = Q := by + rcases Q with ⟨Q⟩ + by_cases hQz : Q z = 0 + · erw [addMap_eq, add_of_Z_eq_zero hP hQ hPz hQz, + smul_eq _ <| (isUnit_Y_of_Z_eq_zero hP hPz).pow 4, Quotient.eq] + exact Setoid.symm <| equiv_zero_of_Z_eq_zero hQ hQz + · erw [addMap_eq, add_of_Z_eq_zero_left hP.left hPz hQz, + smul_eq _ <| ((isUnit_Y_of_Z_eq_zero hP hPz).pow 2).mul <| Ne.isUnit hQz] + rfl + +lemma addMap_of_Z_eq_zero_right {P : PointClass F} {Q : Fin 3 → F} (hP : W.NonsingularLift P) + (hQ : W.Nonsingular Q) (hQz : Q z = 0) : W.addMap P ⟦Q⟧ = P := by + rcases P with ⟨P⟩ + by_cases hPz : P z = 0 + · erw [addMap_eq, add_of_Z_eq_zero hP hQ hPz hQz, + smul_eq _ <| (isUnit_Y_of_Z_eq_zero hP hPz).pow 4, Quotient.eq] + exact Setoid.symm <| equiv_zero_of_Z_eq_zero hP hPz + · erw [addMap_eq, add_of_Z_eq_zero_right hQ.left hPz hQz, + smul_eq _ (((isUnit_Y_of_Z_eq_zero hQ hQz).pow 2).mul <| Ne.isUnit hPz).neg] + rfl + +lemma addMap_of_Y_eq {P Q : Fin 3 → F} (hP : W.Nonsingular P) (hQ : W.Equation Q) (hPz : P z ≠ 0) + (hQz : Q z ≠ 0) (hx : P x * Q z = Q x * P z) (hy' : P y * Q z = W.negY Q * P z) : + W.addMap ⟦P⟧ ⟦Q⟧ = ⟦![0, 1, 0]⟧ := by + by_cases hy : P y * Q z = Q y * P z + · rw [addMap_eq, add_of_Y_eq hP.left hPz hQz hx hy hy', + smul_eq _ <| isUnit_dblU_of_Y_eq hP hPz hQz hx hy hy'] + · rw [addMap_eq, add_of_Y_ne hP.left hQ hPz hQz hx hy, + smul_eq _ <| isUnit_addU_of_Y_ne hPz hQz hy] + +lemma addMap_of_Z_ne_zero {P Q : Fin 3 → F} (hP : W.Equation P) (hQ : W.Equation Q) (hPz : P z ≠ 0) + (hQz : Q z ≠ 0) (hxy : P x * Q z = Q x * P z → P y * Q z ≠ W.negY Q * P z) : W.addMap ⟦P⟧ ⟦Q⟧ = + ⟦![W.toAffine.addX (P x / P z) (Q x / Q z) + (W.toAffine.slope (P x / P z) (Q x / Q z) (P y / P z) (Q y / Q z)), + W.toAffine.addY (P x / P z) (Q x / Q z) (P y / P z) + (W.toAffine.slope (P x / P z) (Q x / Q z) (P y / P z) (Q y / Q z)), 1]⟧ := by + by_cases hx : P x * Q z = Q x * P z + · rw [addMap_eq, add_of_Y_ne' hP hQ hPz hQz hx <| hxy hx, + smul_eq _ <| isUnit_dblZ_of_Y_ne' hP hQ hPz hQz hx <| hxy hx] + · rw [addMap_eq, add_of_X_ne hP hQ hPz hQz hx, smul_eq _ <| isUnit_addZ_of_X_ne hP hQ hx] + +lemma nonsingularLift_addMap {P Q : PointClass F} (hP : W.NonsingularLift P) + (hQ : W.NonsingularLift Q) : W.NonsingularLift <| W.addMap P Q := by + rcases P; rcases Q + exact nonsingular_add hP hQ + +end Addition + end WeierstrassCurve.Projective diff --git a/scripts/no_lints_prime_decls.txt b/scripts/no_lints_prime_decls.txt index 8affabbf01733..671479377b629 100644 --- a/scripts/no_lints_prime_decls.txt +++ b/scripts/no_lints_prime_decls.txt @@ -4796,6 +4796,7 @@ WeierstrassCurve.leadingCoeff_preΨ' WeierstrassCurve.map_preΨ' WeierstrassCurve.natDegree_coeff_preΨ' WeierstrassCurve.natDegree_preΨ' +WeierstrassCurve.Projective.add_of_Y_ne' WeierstrassCurve.Projective.addX_eq' WeierstrassCurve.Projective.addY_of_X_eq' WeierstrassCurve.Projective.addZ_eq' From abad5bb64ad342b9f7f122f073225ba5c551664f Mon Sep 17 00:00:00 2001 From: FR Date: Fri, 18 Oct 2024 18:38:50 +0000 Subject: [PATCH 160/505] chore(Data/Nat/Bits): do not use tactic to make definitions (#17926) --- Mathlib/Data/Nat/Bits.lean | 16 +++++++++------- 1 file changed, 9 insertions(+), 7 deletions(-) diff --git a/Mathlib/Data/Nat/Bits.lean b/Mathlib/Data/Nat/Bits.lean index e580011668c94..63469d742f212 100644 --- a/Mathlib/Data/Nat/Bits.lean +++ b/Mathlib/Data/Nat/Bits.lean @@ -375,10 +375,11 @@ def binaryRec' {motive : ℕ → Sort*} (z : motive 0) ∀ n, motive n := binaryRec z fun b n ih => if h : n = 0 → b = true then f b n h ih - else by - convert z - rw [bit_eq_zero_iff] - simpa using h + else + have : bit b n = 0 := by + rw [bit_eq_zero_iff] + cases n <;> cases b <;> simp at h ⊢ + congrArg motive this ▸ z /-- The same as `binaryRec`, but special casing both 0 and 1 as base cases -/ @[elab_as_elim] @@ -386,9 +387,10 @@ def binaryRecFromOne {motive : ℕ → Sort*} (z₀ : motive 0) (z₁ : motive 1 (f : ∀ b n, n ≠ 0 → motive n → motive (bit b n)) : ∀ n, motive n := binaryRec' z₀ fun b n h ih => - if h' : n = 0 then by - rw [h', h h'] - exact z₁ + if h' : n = 0 then + have : bit b n = bit true 0 := by + rw [h', h h'] + congrArg motive this ▸ z₁ else f b n h' ih @[simp] From 7d20facb45c8983b0c9e7d2db53146a6642b8ef8 Mon Sep 17 00:00:00 2001 From: FordUniver Date: Fri, 18 Oct 2024 19:09:03 +0000 Subject: [PATCH 161/505] feat: variant of the binary AM-GM inequality (#17877) added variant of binary AM-GM inequality Co-authored-by: FordUniver <61389961+FordUniver@users.noreply.github.com> --- Mathlib/Algebra/Order/Ring/Unbundled/Basic.lean | 17 +++++++++++++++-- Mathlib/Data/Nat/Cast/Defs.lean | 4 ++++ 2 files changed, 19 insertions(+), 2 deletions(-) diff --git a/Mathlib/Algebra/Order/Ring/Unbundled/Basic.lean b/Mathlib/Algebra/Order/Ring/Unbundled/Basic.lean index 2aa45b63b34c9..9c0bbff7ea773 100644 --- a/Mathlib/Algebra/Order/Ring/Unbundled/Basic.lean +++ b/Mathlib/Algebra/Order/Ring/Unbundled/Basic.lean @@ -751,8 +751,8 @@ lemma max_mul_mul_le_max_mul_max [PosMulMono α] [MulPosMono α] (b c : α) (ha mul_le_mul (le_max_right a c) (le_max_right b d) hd (le_trans ha (le_max_left a c)) max_le (by simpa [mul_comm, max_comm] using ba) (by simpa [mul_comm, max_comm] using cd) -/-- Binary **arithmetic mean-geometric mean inequality** (aka AM-GM inequality) for linearly ordered -commutative semirings. -/ +/-- Binary, squared, and division-free **arithmetic mean-geometric mean inequality** +(aka AM-GM inequality) for linearly ordered commutative semirings. -/ lemma two_mul_le_add_sq [ExistsAddOfLE α] [MulPosStrictMono α] [ContravariantClass α α (· + ·) (· ≤ ·)] [CovariantClass α α (· + ·) (· ≤ ·)] (a b : α) : 2 * a * b ≤ a ^ 2 + b ^ 2 := by @@ -761,6 +761,19 @@ lemma two_mul_le_add_sq [ExistsAddOfLE α] [MulPosStrictMono α] alias two_mul_le_add_pow_two := two_mul_le_add_sq +/-- Binary, squared, and division-free **arithmetic mean-geometric mean inequality** +(aka AM-GM inequality) for linearly ordered commutative semirings. -/ +lemma four_mul_le_sq_add [ExistsAddOfLE α] [MulPosStrictMono α] + [ContravariantClass α α (· + ·) (· ≤ ·)] [CovariantClass α α (· + ·) (· ≤ ·)] + (a b : α) : 4 * a * b ≤ (a + b) ^ 2 := by + calc 4 * a * b + _ = 2 * a * b + 2 * a * b := by rw [mul_assoc, two_add_two_eq_four.symm, add_mul, mul_assoc] + _ ≤ a ^ 2 + b ^ 2 + 2 * a * b := by gcongr; exact two_mul_le_add_sq _ _ + _ = a ^ 2 + 2 * a * b + b ^ 2 := by rw [add_right_comm] + _ = (a + b) ^ 2 := (add_sq a b).symm + +alias four_mul_le_pow_two_add := four_mul_le_sq_add + end LinearOrderedCommSemiring section LinearOrderedRing diff --git a/Mathlib/Data/Nat/Cast/Defs.lean b/Mathlib/Data/Nat/Cast/Defs.lean index 937ea2f0f088c..eb1e5e1690360 100644 --- a/Mathlib/Data/Nat/Cast/Defs.lean +++ b/Mathlib/Data/Nat/Cast/Defs.lean @@ -199,3 +199,7 @@ theorem three_add_one_eq_four [AddMonoidWithOne R] : 3 + 1 = (4 : R) := by ← Nat.cast_add, ← Nat.cast_add, ← Nat.cast_add] apply congrArg decide + +theorem two_add_two_eq_four [AddMonoidWithOne R] : 2 + 2 = (4 : R) := by + simp [← one_add_one_eq_two, ← Nat.cast_one, ← three_add_one_eq_four, + ← two_add_one_eq_three, add_assoc] From 63c47e1d61626a0e8d3209728b4a86fc90b87cd7 Mon Sep 17 00:00:00 2001 From: FR Date: Fri, 18 Oct 2024 19:09:04 +0000 Subject: [PATCH 162/505] chore(Data/Nat/Bits): attribute `elab_as_elim` `inline` `specialize` (#17897) Add attribute `elab_as_elim` to binary recursors. Also add `inline` and `specialize` for performance. Without `specialize`, the `popcountTR` in the test is not tail-recursive, and so crashes with ``` error: libc++abi: terminating due to uncaught exception of type lean::throwable: deep recursion was detected at 'interpreter' ``` Co-authored-by: Eric Wieser --- Mathlib/Data/Nat/Bits.lean | 10 ++++++---- Mathlib/Data/Nat/Fib/Basic.lean | 8 ++++---- Mathlib/Data/Nat/Size.lean | 14 +++++++------- test/BinaryRec.lean | 8 ++++++++ 4 files changed, 25 insertions(+), 15 deletions(-) create mode 100644 test/BinaryRec.lean diff --git a/Mathlib/Data/Nat/Bits.lean b/Mathlib/Data/Nat/Bits.lean index 63469d742f212..371661d05995b 100644 --- a/Mathlib/Data/Nat/Bits.lean +++ b/Mathlib/Data/Nat/Bits.lean @@ -136,6 +136,7 @@ theorem bit_testBit_zero_shiftRight_one (n : Nat) : bit (n.testBit 0) (n >>> 1) /-- For a predicate `motive : Nat → Sort*`, if instances can be constructed for natural numbers of the form `bit b n`, they can be constructed for any given natural number. -/ +@[inline] def bitCasesOn {motive : Nat → Sort u} (n) (h : ∀ b n, motive (bit b n)) : motive n := -- `1 &&& n != 0` is faster than `n.testBit 0`. This may change when we have faster `testBit`. let x := h (1 &&& n != 0) (n >>> 1) @@ -175,6 +176,7 @@ lemma binaryRec_decreasing (h : n ≠ 0) : div2 n < n := by For a predicate `motive : Nat → Sort*`, if instances can be constructed for natural numbers of the form `bit b n`, they can be constructed for all natural numbers. -/ +@[elab_as_elim, specialize] def binaryRec {motive : Nat → Sort u} (z : motive 0) (f : ∀ b n, motive n → motive (bit b n)) (n : Nat) : motive n := if n0 : n = 0 then congrArg motive n0 ▸ z @@ -208,7 +210,7 @@ lemma binaryRec_zero {motive : Nat → Sort u} @[simp] lemma binaryRec_one {C : Nat → Sort u} (z : C 0) (f : ∀ b n, C n → C (bit b n)) : - binaryRec z f 1 = f true 0 z := by + binaryRec (motive := C) z f 1 = f true 0 z := by rw [binaryRec] simp only [succ_ne_self, ↓reduceDIte, reduceShiftRight, binaryRec_zero] rfl @@ -369,7 +371,7 @@ theorem binaryRec_eq {motive : Nat → Sort u} {z : motive 0} /-- The same as `binaryRec`, but the induction step can assume that if `n=0`, the bit being appended is `true`-/ -@[elab_as_elim] +@[elab_as_elim, specialize] def binaryRec' {motive : ℕ → Sort*} (z : motive 0) (f : ∀ b n, (n = 0 → b = true) → motive n → motive (bit b n)) : ∀ n, motive n := @@ -382,7 +384,7 @@ def binaryRec' {motive : ℕ → Sort*} (z : motive 0) congrArg motive this ▸ z /-- The same as `binaryRec`, but special casing both 0 and 1 as base cases -/ -@[elab_as_elim] +@[elab_as_elim, specialize] def binaryRecFromOne {motive : ℕ → Sort*} (z₀ : motive 0) (z₁ : motive 1) (f : ∀ b n, n ≠ 0 → motive n → motive (bit b n)) : ∀ n, motive n := @@ -399,7 +401,7 @@ theorem zero_bits : bits 0 = [] := by simp [Nat.bits] @[simp] theorem bits_append_bit (n : ℕ) (b : Bool) (hn : n = 0 → b = true) : (bit b n).bits = b :: n.bits := by - rw [Nat.bits, binaryRec_eq'] + rw [Nat.bits, Nat.bits, binaryRec_eq'] simpa @[simp] diff --git a/Mathlib/Data/Nat/Fib/Basic.lean b/Mathlib/Data/Nat/Fib/Basic.lean index 4d5698a009a29..587e609b5ab7b 100644 --- a/Mathlib/Data/Nat/Fib/Basic.lean +++ b/Mathlib/Data/Nat/Fib/Basic.lean @@ -200,12 +200,12 @@ theorem fast_fib_aux_bit_tt (n : ℕ) : · simp theorem fast_fib_aux_eq (n : ℕ) : fastFibAux n = (fib n, fib (n + 1)) := by - apply Nat.binaryRec _ (fun b n' ih => _) n + refine Nat.binaryRec ?_ ?_ n · simp [fastFibAux] · rintro (_|_) n' ih <;> - simp only [fast_fib_aux_bit_ff, fast_fib_aux_bit_tt, congr_arg Prod.fst ih, - congr_arg Prod.snd ih, Prod.mk.inj_iff] <;> - simp [bit, fib_two_mul, fib_two_mul_add_one, fib_two_mul_add_two] + simp only [fast_fib_aux_bit_ff, fast_fib_aux_bit_tt, congr_arg Prod.fst ih, + congr_arg Prod.snd ih, Prod.mk.inj_iff] <;> + simp [bit, fib_two_mul, fib_two_mul_add_one, fib_two_mul_add_two] theorem fast_fib_eq (n : ℕ) : fastFib n = fib n := by rw [fastFib, fast_fib_aux_eq] diff --git a/Mathlib/Data/Nat/Size.lean b/Mathlib/Data/Nat/Size.lean index 5c7eb08714c3c..1fdafff0b8fd0 100644 --- a/Mathlib/Data/Nat/Size.lean +++ b/Mathlib/Data/Nat/Size.lean @@ -39,7 +39,7 @@ theorem size_zero : size 0 = 0 := by simp [size] @[simp] theorem size_bit {b n} (h : bit b n ≠ 0) : size (bit b n) = succ (size n) := by - rw [size] + unfold size conv => lhs rw [binaryRec] @@ -81,7 +81,7 @@ theorem size_shiftLeft {m} (h : m ≠ 0) (n) : size (m <<< n) = size m + n := by theorem lt_size_self (n : ℕ) : n < 2 ^ size n := by rw [← one_shiftLeft] have : ∀ {n}, n = 0 → n < 1 <<< (size n) := by simp - apply binaryRec _ _ n + refine binaryRec ?_ ?_ n · apply this rfl intro b n IH by_cases h : bit b n = 0 @@ -91,11 +91,11 @@ theorem lt_size_self (n : ℕ) : n < 2 ^ size n := by theorem size_le {m n : ℕ} : size m ≤ n ↔ m < 2 ^ n := ⟨fun h => lt_of_lt_of_le (lt_size_self _) (pow_le_pow_of_le_right (by decide) h), by - rw [← one_shiftLeft]; revert n - apply binaryRec _ _ m - · intro n - simp - · intro b m IH n h + rw [← one_shiftLeft] + induction m using binaryRec generalizing n with + | z => simp + | f b m IH => + intro h by_cases e : bit b m = 0 · simp [e] rw [size_bit e] diff --git a/test/BinaryRec.lean b/test/BinaryRec.lean new file mode 100644 index 0000000000000..1112d600d4af8 --- /dev/null +++ b/test/BinaryRec.lean @@ -0,0 +1,8 @@ +import Mathlib.Data.Nat.Bits + +def Nat.popcountTR (n : Nat) : Nat := + n.binaryRec (·) (fun b _ f x ↦ f (x + b.toNat)) 0 + +/-- info: 1 -/ +#guard_msgs in +#eval Nat.popcountTR (2 ^ 20240) From 547e878793a20c0a59b55ba0cca411f38d0e1bfa Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 18 Oct 2024 21:00:38 +0000 Subject: [PATCH 163/505] =?UTF-8?q?feat:=20`Disjoint=20(a=20=E2=80=A2=20s)?= =?UTF-8?q?=20t=20=E2=86=94=20Disjoint=20s=20(a=E2=81=BB=C2=B9=20=E2=80=A2?= =?UTF-8?q?=20t)`=20(#17907)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- Mathlib/Data/Set/Pointwise/SMul.lean | 17 ++++++++++++++--- 1 file changed, 14 insertions(+), 3 deletions(-) diff --git a/Mathlib/Data/Set/Pointwise/SMul.lean b/Mathlib/Data/Set/Pointwise/SMul.lean index 00fe29ebfea63..9e6917fe0172c 100644 --- a/Mathlib/Data/Set/Pointwise/SMul.lean +++ b/Mathlib/Data/Set/Pointwise/SMul.lean @@ -330,7 +330,7 @@ theorem smul_set_inter : a • (s ∩ t) = a • s ∩ a • t := image_inter <| MulAction.injective a @[to_additive] -theorem smul_set_iInter {ι : Type*} +theorem smul_set_iInter {ι : Sort*} (a : α) (t : ι → Set β) : (a • ⋂ i, t i) = ⋂ i, a • t i := image_iInter (MulAction.bijective a) t @@ -401,8 +401,19 @@ lemma inv_op_smul_set_distrib (a : α) (s : Set α) : (op a • s)⁻¹ = a⁻¹ ext; simp [mem_smul_set_iff_inv_smul_mem] @[to_additive (attr := simp)] -lemma smul_set_disjoint_iff : Disjoint (a • s) (a • t) ↔ Disjoint s t := by - simp [disjoint_iff, ← smul_set_inter] +lemma disjoint_smul_set : Disjoint (a • s) (a • t) ↔ Disjoint s t := + disjoint_image_iff <| MulAction.injective _ + +@[to_additive] +lemma disjoint_smul_set_left : Disjoint (a • s) t ↔ Disjoint s (a⁻¹ • t) := by + simpa using disjoint_smul_set (a := a) (t := a⁻¹ • t) + +@[to_additive] +lemma disjoint_smul_set_right : Disjoint s (a • t) ↔ Disjoint (a⁻¹ • s) t := by + simpa using disjoint_smul_set (a := a) (s := a⁻¹ • s) + +@[to_additive (attr := deprecated (since := "2024-10-18"))] +alias smul_set_disjoint_iff := disjoint_smul_set end Group From 8959a60fa36fe31f229b95fa4bb3612b36b29160 Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Fri, 18 Oct 2024 21:29:10 +0000 Subject: [PATCH 164/505] feat(CategoryTheory/Monoidal): add `IsMon_Hom` type class (#17186) --- Mathlib/CategoryTheory/Monoidal/Bimon_.lean | 4 ++++ Mathlib/CategoryTheory/Monoidal/Comon_.lean | 9 +++++++++ Mathlib/CategoryTheory/Monoidal/Mon_.lean | 9 +++++++++ 3 files changed, 22 insertions(+) diff --git a/Mathlib/CategoryTheory/Monoidal/Bimon_.lean b/Mathlib/CategoryTheory/Monoidal/Bimon_.lean index e195b65504704..d522837730846 100644 --- a/Mathlib/CategoryTheory/Monoidal/Bimon_.lean +++ b/Mathlib/CategoryTheory/Monoidal/Bimon_.lean @@ -68,6 +68,10 @@ theorem one_counit (M : C) [Bimon_Class M] : η[M] ≫ ε[M] = 𝟙 (𝟙_ C) := end Bimon_Class +/-- The property that a morphism between bimonoid objects is a bimonoid morphism. -/ +class IsBimon_Hom {M N : C} [Bimon_Class M] [Bimon_Class N] (f : M ⟶ N) extends + IsMon_Hom f, IsComon_Hom f : Prop + variable (C) /-- diff --git a/Mathlib/CategoryTheory/Monoidal/Comon_.lean b/Mathlib/CategoryTheory/Monoidal/Comon_.lean index ff63ce4bbb0eb..ac2283e8e7402 100644 --- a/Mathlib/CategoryTheory/Monoidal/Comon_.lean +++ b/Mathlib/CategoryTheory/Monoidal/Comon_.lean @@ -71,6 +71,15 @@ end Comon_Class open scoped Comon_Class +variable {M N : C} [Comon_Class M] [Comon_Class N] + +/-- The property that a morphism between comonoid objects is a comonoid morphism. -/ +class IsComon_Hom (f : M ⟶ N) : Prop where + hom_counit : f ≫ ε = ε := by aesop_cat + hom_comul : f ≫ Δ = Δ ≫ (f ⊗ f) := by aesop_cat + +attribute [reassoc (attr := simp)] IsComon_Hom.hom_counit IsComon_Hom.hom_comul + variable (C) /-- A comonoid object internal to a monoidal category. diff --git a/Mathlib/CategoryTheory/Monoidal/Mon_.lean b/Mathlib/CategoryTheory/Monoidal/Mon_.lean index c04c4f4d55c07..52b4de6677a39 100644 --- a/Mathlib/CategoryTheory/Monoidal/Mon_.lean +++ b/Mathlib/CategoryTheory/Monoidal/Mon_.lean @@ -67,6 +67,15 @@ end Mon_Class open scoped Mon_Class +variable {M N : C} [Mon_Class M] [Mon_Class N] + +/-- The property that a morphism between monoid objects is a monoid morphism. -/ +class IsMon_Hom (f : M ⟶ N) : Prop where + one_hom : η ≫ f = η := by aesop_cat + mul_hom : μ ≫ f = (f ⊗ f) ≫ μ := by aesop_cat + +attribute [reassoc (attr := simp)] IsMon_Hom.one_hom IsMon_Hom.mul_hom + variable (C) /-- A monoid object internal to a monoidal category. From de5d061f84a1fdff844f983aea367c3bdbf676b8 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 18 Oct 2024 21:29:11 +0000 Subject: [PATCH 165/505] chore: Rename `OpenEmbedding` to `IsOpenEmbedding` (#17898) `Function.Embedding` is a type while `Embedding` is a proposition, and there are many other kinds of embeddings than topological embeddings. Hence this PR is a step towards 1. renaming `Embedding` to `IsEmbedding` and similarly for neighborhing declarations (which `OpenEmbedding` is) 2. namespacing it inside `Topology` [Zulip](https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/rename.20.60Inducing.60.20and.20.60Embedding.60.3F). See #15993 for context. --- Mathlib/AlgebraicGeometry/AffineScheme.lean | 14 +- Mathlib/AlgebraicGeometry/Cover/Open.lean | 2 +- Mathlib/AlgebraicGeometry/FunctionField.lean | 4 +- Mathlib/AlgebraicGeometry/Gluing.lean | 9 +- Mathlib/AlgebraicGeometry/Limits.lean | 4 +- .../AlgebraicGeometry/Morphisms/Affine.lean | 2 +- .../AlgebraicGeometry/Morphisms/IsIso.lean | 2 +- .../Morphisms/OpenImmersion.lean | 4 +- .../Morphisms/Preimmersion.lean | 2 +- .../Morphisms/UnderlyingMap.lean | 14 +- Mathlib/AlgebraicGeometry/Noetherian.lean | 4 +- Mathlib/AlgebraicGeometry/OpenImmersion.lean | 25 ++-- .../PrimeSpectrum/Basic.lean | 7 +- .../ProjectiveSpectrum/Scheme.lean | 13 +- Mathlib/AlgebraicGeometry/Properties.lean | 4 +- Mathlib/AlgebraicGeometry/Restrict.lean | 30 ++--- Mathlib/AlgebraicGeometry/Scheme.lean | 2 +- .../Complex/UpperHalfPlane/Manifold.lean | 4 +- .../Complex/UpperHalfPlane/Topology.lean | 15 ++- Mathlib/Analysis/Normed/Ring/Units.lean | 7 +- Mathlib/Analysis/SpecialFunctions/Exp.lean | 11 +- Mathlib/Geometry/Manifold/ChartedSpace.lean | 12 +- .../Geometry/Manifold/ContMDiff/Basic.lean | 27 ++-- .../Instances/UnitsOfNormedAlgebra.lean | 10 +- .../Manifold/LocalInvariantProperties.lean | 4 +- .../Manifold/SmoothManifoldWithCorners.lean | 8 +- .../RingedSpace/LocallyRingedSpace.lean | 10 +- .../Geometry/RingedSpace/OpenImmersion.lean | 53 ++++---- .../Geometry/RingedSpace/PresheafedSpace.lean | 16 +-- .../RingedSpace/PresheafedSpace/Gluing.lean | 19 +-- .../Geometry/RingedSpace/SheafedSpace.lean | 8 +- Mathlib/Geometry/RingedSpace/Stalks.lean | 10 +- .../Constructions/BorelSpace/Basic.lean | 5 +- .../EisensteinSeries/UniformConvergence.lean | 2 +- Mathlib/Topology/Algebra/ConstMulAction.lean | 2 +- .../Algebra/Constructions/DomMulAct.lean | 12 +- .../Algebra/Nonarchimedean/Basic.lean | 2 +- Mathlib/Topology/Bases.lean | 4 +- .../Category/CompHausLike/Limits.lean | 28 ++-- Mathlib/Topology/Category/Stonean/Limits.lean | 4 +- Mathlib/Topology/Category/TopCat/Basic.lean | 38 ++++-- .../Category/TopCat/Limits/Products.lean | 7 +- .../Category/TopCat/Limits/Pullbacks.lean | 48 ++++--- .../Topology/Category/TopCat/OpenNhds.lean | 7 +- Mathlib/Topology/Category/TopCat/Opens.lean | 33 +++-- Mathlib/Topology/Clopen.lean | 2 +- .../Topology/Compactification/OnePoint.lean | 23 ++-- .../Topology/Compactness/LocallyCompact.lean | 7 +- .../Topology/Connected/LocallyConnected.lean | 9 +- Mathlib/Topology/Connected/PathConnected.lean | 7 +- Mathlib/Topology/Constructions.lean | 72 +++++++---- Mathlib/Topology/ContinuousOn.lean | 5 +- Mathlib/Topology/Defs/Induced.lean | 7 +- Mathlib/Topology/Gluing.lean | 24 ++-- Mathlib/Topology/Homeomorph.lean | 16 ++- Mathlib/Topology/Inseparable.lean | 5 +- Mathlib/Topology/Instances/ENNReal.lean | 11 +- Mathlib/Topology/Instances/ENat.lean | 7 +- Mathlib/Topology/Instances/EReal.lean | 9 +- Mathlib/Topology/Irreducible.lean | 2 +- Mathlib/Topology/IsLocalHomeomorph.lean | 54 +++++--- Mathlib/Topology/LocalAtTarget.lean | 30 +++-- Mathlib/Topology/LocallyClosed.lean | 5 +- Mathlib/Topology/Maps/Basic.lean | 121 ++++++++++++------ Mathlib/Topology/PartialHomeomorph.lean | 29 +++-- Mathlib/Topology/QuasiSeparated.lean | 14 +- Mathlib/Topology/SeparatedMap.lean | 11 +- Mathlib/Topology/Separation.lean | 17 ++- Mathlib/Topology/Sets/Opens.lean | 14 +- Mathlib/Topology/Sheaves/LocalPredicate.lean | 4 +- .../Sheaves/SheafCondition/Sites.lean | 19 ++- Mathlib/Topology/Sheaves/Stalks.lean | 7 +- Mathlib/Topology/Sober.lean | 7 +- Mathlib/Topology/Support.lean | 2 +- scripts/no_lints_prime_decls.txt | 8 +- 75 files changed, 677 insertions(+), 419 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/AffineScheme.lean b/Mathlib/AlgebraicGeometry/AffineScheme.lean index 74ea6759b202f..5a69b75178c41 100644 --- a/Mathlib/AlgebraicGeometry/AffineScheme.lean +++ b/Mathlib/AlgebraicGeometry/AffineScheme.lean @@ -278,7 +278,7 @@ def isoSpec : ↑U ≅ Spec Γ(X, U) := haveI : IsAffine U := hU U.toScheme.isoSpec ≪≫ Scheme.Spec.mapIso - (X.presheaf.mapIso (eqToIso U.openEmbedding_obj_top).op).op + (X.presheaf.mapIso (eqToIso U.isOpenEmbedding_obj_top).op).op open LocalRing in lemma isoSpec_hom_base_apply (x : U) : @@ -395,8 +395,8 @@ theorem _root_.AlgebraicGeometry.Scheme.Hom.isAffineOpen_iff_of_isOpenImmersion (f : AlgebraicGeometry.Scheme.Hom X Y) [H : IsOpenImmersion f] {U : X.Opens} : IsAffineOpen (f ''ᵁ U) ↔ IsAffineOpen U := by refine ⟨fun hU => @isAffine_of_isIso _ _ - (IsOpenImmersion.isoOfRangeEq (X.ofRestrict U.openEmbedding ≫ f) (Y.ofRestrict _) ?_).hom ?_ hU, - fun hU => hU.image_of_isOpenImmersion f⟩ + (IsOpenImmersion.isoOfRangeEq (X.ofRestrict U.isOpenEmbedding ≫ f) (Y.ofRestrict _) ?_).hom + ?_ hU, fun hU => hU.image_of_isOpenImmersion f⟩ · rw [Scheme.comp_base, coe_comp, Set.range_comp] dsimp [Opens.coe_inclusion', Scheme.restrict] erw [Subtype.range_coe, Subtype.range_coe] -- now `erw` after #13170 @@ -491,7 +491,7 @@ theorem ι_basicOpen_preimage (r : Γ(X, ⊤)) : IsAffineOpen ((X.basicOpen r).ι ⁻¹ᵁ U) := by apply (X.basicOpen r).ι.isAffineOpen_iff_of_isOpenImmersion.mp dsimp [Scheme.Hom.opensFunctor, LocallyRingedSpace.IsOpenImmersion.opensFunctor] - rw [Opens.functor_obj_map_obj, Opens.openEmbedding_obj_top, inf_comm, + rw [Opens.functor_obj_map_obj, Opens.isOpenEmbedding_obj_top, inf_comm, ← Scheme.basicOpen_res _ _ (homOfLE le_top).op] exact hU.basicOpen _ @@ -504,12 +504,12 @@ theorem exists_basicOpen_le {V : X.Opens} (x : V) (h : ↑x ∈ U) : ((Opens.map U.inclusion').obj V).isOpen have : U.ι ''ᵁ (U.toScheme.basicOpen r) = - X.basicOpen (X.presheaf.map (eqToHom U.openEmbedding_obj_top.symm).op r) := by + X.basicOpen (X.presheaf.map (eqToHom U.isOpenEmbedding_obj_top.symm).op r) := by refine (Scheme.image_basicOpen U.ι r).trans ?_ rw [Scheme.basicOpen_res_eq] simp only [Scheme.Opens.toScheme_presheaf_obj, Scheme.Opens.ι_appIso, Iso.refl_inv, CommRingCat.id_apply] - use X.presheaf.map (eqToHom U.openEmbedding_obj_top.symm).op r + use X.presheaf.map (eqToHom U.isOpenEmbedding_obj_top.symm).op r rw [← this] exact ⟨Set.image_subset_iff.mpr h₂, ⟨_, h⟩, h₁, rfl⟩ @@ -573,7 +573,7 @@ theorem isLocalization_of_eq_basicOpen {V : X.Opens} (i : V ⟶ U) (e : V = X.ba instance _root_.AlgebraicGeometry.Γ_restrict_isLocalization (X : Scheme.{u}) [IsAffine X] (r : Γ(X, ⊤)) : IsLocalization.Away r Γ(X.basicOpen r, ⊤) := - (isAffineOpen_top X).isLocalization_of_eq_basicOpen r _ (Opens.openEmbedding_obj_top _) + (isAffineOpen_top X).isLocalization_of_eq_basicOpen r _ (Opens.isOpenEmbedding_obj_top _) include hU in theorem basicOpen_basicOpen_is_basicOpen (g : Γ(X, X.basicOpen f)) : diff --git a/Mathlib/AlgebraicGeometry/Cover/Open.lean b/Mathlib/AlgebraicGeometry/Cover/Open.lean index 836bfdf366a33..32b601db856c3 100644 --- a/Mathlib/AlgebraicGeometry/Cover/Open.lean +++ b/Mathlib/AlgebraicGeometry/Cover/Open.lean @@ -248,7 +248,7 @@ theorem OpenCover.compactSpace {X : Scheme.{u}} (𝒰 : X.OpenCover) [Finite (TopCat.homeoOfIso (asIso (IsOpenImmersion.isoOfRangeEq (𝒰.map i) - (X.ofRestrict (Opens.openEmbedding ⟨_, (𝒰.IsOpen i).base_open.isOpen_range⟩)) + (X.ofRestrict (Opens.isOpenEmbedding ⟨_, (𝒰.IsOpen i).base_open.isOpen_range⟩)) Subtype.range_coe.symm).hom.base)) /-- Given open covers `{ Uᵢ }` and `{ Uⱼ }`, we may form the open cover `{ Uᵢ ∩ Uⱼ }`. -/ diff --git a/Mathlib/AlgebraicGeometry/FunctionField.lean b/Mathlib/AlgebraicGeometry/FunctionField.lean index 242c360e3de84..6833058adc17f 100644 --- a/Mathlib/AlgebraicGeometry/FunctionField.lean +++ b/Mathlib/AlgebraicGeometry/FunctionField.lean @@ -135,7 +135,7 @@ theorem IsAffineOpen.primeIdealOf_genericPoint {X : Scheme} [IsIntegral X] {U : delta IsAffineOpen.primeIdealOf convert genericPoint_eq_of_isOpenImmersion - (U.toScheme.isoSpec.hom ≫ Spec.map (X.presheaf.map (eqToHom U.openEmbedding_obj_top).op)) + (U.toScheme.isoSpec.hom ≫ Spec.map (X.presheaf.map (eqToHom U.isOpenEmbedding_obj_top).op)) -- Porting note: this was `ext1` apply Subtype.ext exact (genericPoint_eq_of_isOpenImmersion U.ι).symm @@ -146,7 +146,7 @@ theorem functionField_isFractionRing_of_isAffineOpen [IsIntegral X] (U : X.Opens haveI : IsAffine _ := hU haveI : IsIntegral U := @isIntegral_of_isAffine_of_isDomain _ _ _ - (by rw [Scheme.Opens.toScheme_presheaf_obj, Opens.openEmbedding_obj_top]; infer_instance) + (by rw [Scheme.Opens.toScheme_presheaf_obj, Opens.isOpenEmbedding_obj_top]; infer_instance) delta IsFractionRing Scheme.functionField convert hU.isLocalization_stalk ⟨genericPoint X, (((genericPoint_spec X).mem_open_set_iff U.isOpen).mpr (by simpa using ‹Nonempty U›))⟩ using 1 diff --git a/Mathlib/AlgebraicGeometry/Gluing.lean b/Mathlib/AlgebraicGeometry/Gluing.lean index 9eeee3d9e71e3..c38450ee48155 100644 --- a/Mathlib/AlgebraicGeometry/Gluing.lean +++ b/Mathlib/AlgebraicGeometry/Gluing.lean @@ -379,10 +379,13 @@ theorem fromGlued_open_map : IsOpenMap 𝒰.fromGlued.base := by exact Set.preimage_image_eq _ 𝒰.fromGlued_injective · exact ⟨hx, 𝒰.covers x⟩ -theorem fromGlued_openEmbedding : OpenEmbedding 𝒰.fromGlued.base := - openEmbedding_of_continuous_injective_open +theorem fromGlued_isOpenEmbedding : IsOpenEmbedding 𝒰.fromGlued.base := + isOpenEmbedding_of_continuous_injective_open (by fun_prop) 𝒰.fromGlued_injective 𝒰.fromGlued_open_map +@[deprecated (since := "2024-10-18")] +alias fromGlued_openEmbedding := fromGlued_isOpenEmbedding + instance : Epi 𝒰.fromGlued.base := by rw [TopCat.epi_iff_surjective] intro x @@ -393,7 +396,7 @@ instance : Epi 𝒰.fromGlued.base := by exact h instance fromGlued_open_immersion : IsOpenImmersion 𝒰.fromGlued := - IsOpenImmersion.of_stalk_iso _ 𝒰.fromGlued_openEmbedding + IsOpenImmersion.of_stalk_iso _ 𝒰.fromGlued_isOpenEmbedding instance : IsIso 𝒰.fromGlued := let F := Scheme.forgetToLocallyRingedSpace ⋙ LocallyRingedSpace.forgetToSheafedSpace ⋙ diff --git a/Mathlib/AlgebraicGeometry/Limits.lean b/Mathlib/AlgebraicGeometry/Limits.lean index 0622355eef447..1480c9d5d315f 100644 --- a/Mathlib/AlgebraicGeometry/Limits.lean +++ b/Mathlib/AlgebraicGeometry/Limits.lean @@ -250,7 +250,7 @@ lemma sigmaι_eq_iff (i j : ι) (x y) : by_cases h : i = j · subst h simp only [Sigma.mk.inj_iff, heq_eq_eq, true_and] - exact ((disjointGlueData f).ι i).openEmbedding.inj H + exact ((disjointGlueData f).ι i).isOpenEmbedding.inj H · obtain (e | ⟨z, _⟩) := (Scheme.GlueData.ι_eq_iff _ _ _ _ _).mp H · exact (h (Sigma.mk.inj_iff.mp e).1).elim · simp only [disjointGlueData_J, disjointGlueData_V, h, ↓reduceIte] at z @@ -269,7 +269,7 @@ lemma disjoint_opensRange_sigmaι (i j : ι) (h : i ≠ j) : lemma exists_sigmaι_eq (x : (∐ f : _)) : ∃ i y, (Sigma.ι f i).base y = x := by obtain ⟨i, y, e⟩ := (disjointGlueData f).ι_jointly_surjective ((sigmaIsoGlued f).hom.base x) - refine ⟨i, y, (sigmaIsoGlued f).hom.openEmbedding.inj ?_⟩ + refine ⟨i, y, (sigmaIsoGlued f).hom.isOpenEmbedding.inj ?_⟩ rwa [← Scheme.comp_base_apply, ι_sigmaIsoGlued_hom] lemma iSup_opensRange_sigmaι : ⨆ i, (Sigma.ι f i).opensRange = ⊤ := diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Affine.lean b/Mathlib/AlgebraicGeometry/Morphisms/Affine.lean index ecd431a4735b4..bcb0509ca655d 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Affine.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Affine.lean @@ -129,7 +129,7 @@ lemma isAffine_of_isAffineOpen_basicOpen (s : Set Γ(X, ⊤)) refine IsIso.comp_isIso' ?_ inferInstance convert isIso_ΓSpec_adjunction_unit_app_basicOpen i.1 using 0 refine congr(IsIso ((ΓSpec.adjunction.unit.app X).app $(?_))) - rw [Opens.openEmbedding_obj_top] + rw [Opens.isOpenEmbedding_obj_top] /-- If `s` is a spanning set of `Γ(X, U)`, such that each `X.basicOpen i` is affine, then `U` is also diff --git a/Mathlib/AlgebraicGeometry/Morphisms/IsIso.lean b/Mathlib/AlgebraicGeometry/Morphisms/IsIso.lean index 384f3493f9a1e..41f31c85bebb9 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/IsIso.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/IsIso.lean @@ -31,7 +31,7 @@ lemma isomorphisms_eq_stalkwise : ext X Y f exact ⟨fun H ↦ inferInstanceAs (IsIso (TopCat.isoOfHomeo (H.1.1.toHomeomeomorph_of_surjective H.2)).hom), fun (_ : IsIso f.base) ↦ - let e := (TopCat.homeoOfIso <| asIso f.base); ⟨e.openEmbedding, e.surjective⟩⟩ + let e := (TopCat.homeoOfIso <| asIso f.base); ⟨e.isOpenEmbedding, e.surjective⟩⟩ instance : IsLocalAtTarget (isomorphisms Scheme) := isomorphisms_eq_isOpenImmersion_inf_surjective ▸ inferInstance diff --git a/Mathlib/AlgebraicGeometry/Morphisms/OpenImmersion.lean b/Mathlib/AlgebraicGeometry/Morphisms/OpenImmersion.lean index 43a94b8d2d1a0..1a4868bc3510b 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/OpenImmersion.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/OpenImmersion.lean @@ -29,13 +29,13 @@ namespace AlgebraicGeometry variable {X Y Z : Scheme.{u}} (f : X ⟶ Y) (g : Y ⟶ Z) theorem isOpenImmersion_iff_stalk {f : X ⟶ Y} : IsOpenImmersion f ↔ - OpenEmbedding f.base ∧ ∀ x, IsIso (f.stalkMap x) := by + IsOpenEmbedding f.base ∧ ∀ x, IsIso (f.stalkMap x) := by constructor · intro h; exact ⟨h.1, inferInstance⟩ · rintro ⟨h₁, h₂⟩; exact IsOpenImmersion.of_stalk_iso f h₁ theorem isOpenImmersion_eq_inf : - @IsOpenImmersion = (topologically OpenEmbedding) ⊓ + @IsOpenImmersion = (topologically IsOpenEmbedding) ⊓ stalkwise (fun f ↦ Function.Bijective f) := by ext exact isOpenImmersion_iff_stalk.trans diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Preimmersion.lean b/Mathlib/AlgebraicGeometry/Morphisms/Preimmersion.lean index a5fd6903836cb..9336dbb858cde 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Preimmersion.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Preimmersion.lean @@ -60,7 +60,7 @@ instance : IsLocalAtTarget @IsPreimmersion := isPreimmersion_eq_inf ▸ inferInstance instance (priority := 900) {X Y : Scheme} (f : X ⟶ Y) [IsOpenImmersion f] : IsPreimmersion f where - base_embedding := f.openEmbedding.toEmbedding + base_embedding := f.isOpenEmbedding.toEmbedding surj_on_stalks _ := (ConcreteCategory.bijective_of_isIso _).2 instance : MorphismProperty.IsMultiplicative @IsPreimmersion where diff --git a/Mathlib/AlgebraicGeometry/Morphisms/UnderlyingMap.lean b/Mathlib/AlgebraicGeometry/Morphisms/UnderlyingMap.lean index ff349f67ab6a8..00d8b44dc43f0 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/UnderlyingMap.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/UnderlyingMap.lean @@ -18,7 +18,7 @@ of the underlying map of topological spaces, including - `IsOpenMap` - `IsClosedMap` - `Embedding` -- `OpenEmbedding` +- `IsOpenEmbedding` - `ClosedEmbedding` -/ @@ -119,17 +119,17 @@ instance embedding_isLocalAtTarget : IsLocalAtTarget (topologically Embedding) : end Embedding -section OpenEmbedding +section IsOpenEmbedding -instance : (topologically OpenEmbedding).RespectsIso := - topologically_respectsIso _ (fun e ↦ e.openEmbedding) (fun _ _ hf hg ↦ hg.comp hf) +instance : (topologically IsOpenEmbedding).RespectsIso := + topologically_respectsIso _ (fun e ↦ e.isOpenEmbedding) (fun _ _ hf hg ↦ hg.comp hf) -instance openEmbedding_isLocalAtTarget : IsLocalAtTarget (topologically OpenEmbedding) := +instance isOpenEmbedding_isLocalAtTarget : IsLocalAtTarget (topologically IsOpenEmbedding) := topologically_isLocalAtTarget _ (fun _ s hf ↦ hf.restrictPreimage s) - (fun _ _ _ hU hfcont hf ↦ (openEmbedding_iff_openEmbedding_of_iSup_eq_top hU hfcont).mpr hf) + (fun _ _ _ hU hfcont hf ↦ (isOpenEmbedding_iff_isOpenEmbedding_of_iSup_eq_top hU hfcont).mpr hf) -end OpenEmbedding +end IsOpenEmbedding section ClosedEmbedding diff --git a/Mathlib/AlgebraicGeometry/Noetherian.lean b/Mathlib/AlgebraicGeometry/Noetherian.lean index 45c741d358319..121bca1ca0d2b 100644 --- a/Mathlib/AlgebraicGeometry/Noetherian.lean +++ b/Mathlib/AlgebraicGeometry/Noetherian.lean @@ -199,7 +199,7 @@ instance (priority := 100) {Z : Scheme} [IsLocallyNoetherian X] apply (quasiCompact_iff_forall_affine f).mpr intro U hU rw [Opens.map_coe, ← Set.preimage_inter_range] - apply f.openEmbedding.toInducing.isCompact_preimage' + apply f.isOpenEmbedding.toInducing.isCompact_preimage' · apply (noetherianSpace_set_iff _).mp · convert noetherianSpace_of_isAffineOpen U hU apply IsLocallyNoetherian.component_noetherian ⟨U, hU⟩ @@ -213,7 +213,7 @@ instance (priority := 100) IsLocallyNoetherian.quasiSeparatedSpace [IsLocallyNoe QuasiSeparatedSpace X := by apply (quasiSeparatedSpace_iff_affine X).mpr intro U V - have hInd := U.2.fromSpec.openEmbedding.toInducing + have hInd := U.2.fromSpec.isOpenEmbedding.toInducing apply (hInd.isCompact_preimage_iff ?_).mp · rw [← Set.preimage_inter_range, IsAffineOpen.range_fromSpec, Set.inter_comm] apply hInd.isCompact_preimage' diff --git a/Mathlib/AlgebraicGeometry/OpenImmersion.lean b/Mathlib/AlgebraicGeometry/OpenImmersion.lean index 7b3769c508404..bc36b58d339d1 100644 --- a/Mathlib/AlgebraicGeometry/OpenImmersion.lean +++ b/Mathlib/AlgebraicGeometry/OpenImmersion.lean @@ -56,7 +56,7 @@ protected def scheme (X : LocallyRingedSpace.{u}) refine SheafedSpace.forgetToPresheafedSpace.preimageIso ?_ apply PresheafedSpace.IsOpenImmersion.isoOfRangeEq (PresheafedSpace.ofRestrict _ _) f.1 · exact Subtype.range_coe_subtype - · exact Opens.openEmbedding _ -- Porting note (#11187): was `infer_instance` + · exact Opens.isOpenEmbedding _ -- Porting note (#11187): was `infer_instance` end LocallyRingedSpace.IsOpenImmersion @@ -71,13 +71,16 @@ namespace Scheme.Hom variable {X Y : Scheme.{u}} (f : Scheme.Hom X Y) [H : IsOpenImmersion f] -theorem openEmbedding : OpenEmbedding f.base := +theorem isOpenEmbedding : IsOpenEmbedding f.base := H.base_open +@[deprecated (since := "2024-10-18")] +alias openEmbedding := isOpenEmbedding + /-- The image of an open immersion as an open set. -/ @[simps] def opensRange : Y.Opens := - ⟨_, f.openEmbedding.isOpen_range⟩ + ⟨_, f.isOpenEmbedding.isOpen_range⟩ /-- The functor `opens X ⥤ opens Y` associated with an open immersion `f : X ⟶ Y`. -/ abbrev opensFunctor : X.Opens ⥤ Y.Opens := @@ -103,7 +106,7 @@ lemma image_top_eq_opensRange : f ''ᵁ ⊤ = f.opensRange := by @[simp] lemma preimage_image_eq (U : X.Opens) : f ⁻¹ᵁ f ''ᵁ U = U := by apply Opens.ext - simp [Set.preimage_image_eq _ f.openEmbedding.inj] + simp [Set.preimage_image_eq _ f.isOpenEmbedding.inj] lemma image_le_image_iff (f : X ⟶ Y) [IsOpenImmersion f] (U U' : X.Opens) : f ''ᵁ U ≤ f ''ᵁ U' ↔ U ≤ U' := by @@ -194,7 +197,7 @@ def IsOpenImmersion.opensEquiv {X Y : Scheme.{u}} (f : X ⟶ Y) [IsOpenImmersion X.Opens ≃ { U : Y.Opens // U ≤ f.opensRange } where toFun U := ⟨f ''ᵁ U, Set.image_subset_range _ _⟩ invFun U := f ⁻¹ᵁ U - left_inv _ := Opens.ext (Set.preimage_image_eq _ f.openEmbedding.inj) + left_inv _ := Opens.ext (Set.preimage_image_eq _ f.isOpenEmbedding.inj) right_inv U := Subtype.ext (Opens.ext (Set.image_preimage_eq_of_subset U.2)) namespace Scheme @@ -202,7 +205,7 @@ namespace Scheme instance basic_open_isOpenImmersion {R : CommRingCat.{u}} (f : R) : IsOpenImmersion (Spec.map (CommRingCat.ofHom (algebraMap R (Localization.Away f)))) := by apply SheafedSpace.IsOpenImmersion.of_stalk_iso (H := ?_) - · exact (PrimeSpectrum.localization_away_openEmbedding (Localization.Away f) f : _) + · exact (PrimeSpectrum.localization_away_isOpenEmbedding (Localization.Away f) f : _) · intro x exact Spec_map_localization_isIso R (Submonoid.powers f) x @@ -292,7 +295,7 @@ end PresheafedSpace.IsOpenImmersion section Restrict -variable {U : TopCat.{u}} (X : Scheme.{u}) {f : U ⟶ TopCat.of X} (h : OpenEmbedding f) +variable {U : TopCat.{u}} (X : Scheme.{u}) {f : U ⟶ TopCat.of X} (h : IsOpenEmbedding f) /-- The restriction of a Scheme along an open embedding. -/ @[simps! (config := .lemmasOnly) carrier, simps! presheaf_obj] @@ -351,7 +354,7 @@ theorem to_iso {X Y : Scheme.{u}} (f : X ⟶ Y) [h : IsOpenImmersion f] [Epi f.b LocallyRingedSpace.forgetToSheafedSpace ⋙ SheafedSpace.forgetToPresheafedSpace) (@PresheafedSpace.IsOpenImmersion.to_iso _ _ _ _ f.toPshHom h _) _ -theorem of_stalk_iso {X Y : Scheme.{u}} (f : X ⟶ Y) (hf : OpenEmbedding f.base) +theorem of_stalk_iso {X Y : Scheme.{u}} (f : X ⟶ Y) (hf : IsOpenEmbedding f.base) [∀ x, IsIso (f.stalkMap x)] : IsOpenImmersion f := haveI (x : X) : IsIso (f.toShHom.stalkMap x) := inferInstanceAs <| IsIso (f.stalkMap x) SheafedSpace.IsOpenImmersion.of_stalk_iso f.toShHom hf @@ -368,10 +371,10 @@ lemma of_comp {X Y Z : Scheme.{u}} (f : X ⟶ Y) (g : Y ⟶ Z) [IsOpenImmersion infer_instance IsIso.of_isIso_comp_left (f := g.stalkMap (f.base x)) _ IsOpenImmersion.of_stalk_iso _ <| - OpenEmbedding.of_comp _ (Scheme.Hom.openEmbedding g) (Scheme.Hom.openEmbedding (f ≫ g)) + IsOpenEmbedding.of_comp _ (Scheme.Hom.isOpenEmbedding g) (Scheme.Hom.isOpenEmbedding (f ≫ g)) theorem iff_stalk_iso {X Y : Scheme.{u}} (f : X ⟶ Y) : - IsOpenImmersion f ↔ OpenEmbedding f.base ∧ ∀ x, IsIso (f.stalkMap x) := + IsOpenImmersion f ↔ IsOpenEmbedding f.base ∧ ∀ x, IsIso (f.stalkMap x) := ⟨fun H => ⟨H.1, fun x ↦ inferInstanceAs <| IsIso (f.toPshHom.stalkMap x)⟩, fun ⟨h₁, h₂⟩ => @IsOpenImmersion.of_stalk_iso _ _ f h₁ h₂⟩ @@ -391,7 +394,7 @@ theorem _root_.AlgebraicGeometry.isIso_iff_stalk_iso {X Y : Scheme.{u}} (f : X (Equiv.ofBijective _ ⟨h₂.inj, (TopCat.epi_iff_surjective _).mp h₁⟩) h₂.continuous h₂.isOpenMap)).hom infer_instance - · intro H; exact ⟨inferInstance, (TopCat.homeoOfIso (asIso f.base)).openEmbedding⟩ + · intro H; exact ⟨inferInstance, (TopCat.homeoOfIso (asIso f.base)).isOpenEmbedding⟩ /-- An open immersion induces an isomorphism from the domain onto the image -/ def isoRestrict : X ≅ (Z.restrict H.base_open : _) := diff --git a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean index 05c884eb52c60..f918a39487b28 100644 --- a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean +++ b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean @@ -457,13 +457,16 @@ theorem localization_away_comap_range (S : Type v) [CommSemiring S] [Algebra R S · rintro h₁ _ ⟨⟨n, rfl⟩, h₃⟩ exact h₁ (x.2.mem_of_pow_mem _ h₃) -theorem localization_away_openEmbedding (S : Type v) [CommSemiring S] [Algebra R S] (r : R) - [IsLocalization.Away r S] : OpenEmbedding (comap (algebraMap R S)) := +theorem localization_away_isOpenEmbedding (S : Type v) [CommSemiring S] [Algebra R S] (r : R) + [IsLocalization.Away r S] : IsOpenEmbedding (comap (algebraMap R S)) := { toEmbedding := localization_comap_embedding S (Submonoid.powers r) isOpen_range := by rw [localization_away_comap_range S r] exact isOpen_basicOpen } +@[deprecated (since := "2024-10-18")] +alias localization_away_openEmbedding := localization_away_isOpenEmbedding + theorem isCompact_basicOpen (f : R) : IsCompact (basicOpen f : Set (PrimeSpectrum R)) := by rw [← localization_away_comap_range (Localization (Submonoid.powers f))] exact isCompact_range (map_continuous _) diff --git a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean index 77da59dbcba24..dae30030d3507 100644 --- a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean +++ b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean @@ -122,12 +122,13 @@ local notation3 "Proj.T" => PresheafedSpace.carrier <| SheafedSpace.toPresheafed /-- `Proj` restrict to some open set -/ macro "Proj| " U:term : term => - `((Proj.toLocallyRingedSpace 𝒜).restrict (Opens.openEmbedding (X := Proj.T) ($U : Opens Proj.T))) + `((Proj.toLocallyRingedSpace 𝒜).restrict + (Opens.isOpenEmbedding (X := Proj.T) ($U : Opens Proj.T))) /-- the underlying topological space of `Proj` restricted to some open set -/ local notation "Proj.T| " U => PresheafedSpace.carrier <| SheafedSpace.toPresheafedSpace <| LocallyRingedSpace.toSheafedSpace - <| (LocallyRingedSpace.restrict Proj (Opens.openEmbedding (X := Proj.T) (U : Opens Proj.T))) + <| (LocallyRingedSpace.restrict Proj (Opens.isOpenEmbedding (X := Proj.T) (U : Opens Proj.T))) /-- basic open sets in `Proj` -/ local notation "pbo " x => ProjectiveSpectrum.basicOpen 𝒜 x @@ -610,13 +611,13 @@ Mathematically, the map is the same as `awayToSection`. -/ def awayToΓ (f) : CommRingCat.of (A⁰_ f) ⟶ LocallyRingedSpace.Γ.obj (op <| Proj| pbo f) := awayToSection 𝒜 f ≫ (ProjectiveSpectrum.Proj.structureSheaf 𝒜).1.map - (homOfLE (Opens.openEmbedding_obj_top _).le).op + (homOfLE (Opens.isOpenEmbedding_obj_top _).le).op lemma awayToΓ_ΓToStalk (f) (x) : awayToΓ 𝒜 f ≫ (Proj| pbo f).presheaf.Γgerm x = HomogeneousLocalization.mapId 𝒜 (Submonoid.powers_le.mpr x.2) ≫ (Proj.stalkIso' 𝒜 x.1).toCommRingCatIso.inv ≫ - ((Proj.toLocallyRingedSpace 𝒜).restrictStalkIso (Opens.openEmbedding _) x).inv := by + ((Proj.toLocallyRingedSpace 𝒜).restrictStalkIso (Opens.isOpenEmbedding _) x).inv := by rw [awayToΓ, Category.assoc, ← Category.assoc _ (Iso.inv _), Iso.eq_comp_inv, Category.assoc, Category.assoc, Presheaf.Γgerm] rw [LocallyRingedSpace.restrictStalkIso_hom_eq_germ] @@ -781,7 +782,7 @@ lemma toStalk_specStalkEquiv (f) (x : pbo f) {m} (f_deg : f ∈ 𝒜 m) (hm : 0 lemma stalkMap_toSpec (f) (x : pbo f) {m} (f_deg : f ∈ 𝒜 m) (hm : 0 < m) : (toSpec 𝒜 f).stalkMap x = (specStalkEquiv 𝒜 f x f_deg hm).hom ≫ (Proj.stalkIso' 𝒜 x.1).toCommRingCatIso.inv ≫ - ((Proj.toLocallyRingedSpace 𝒜).restrictStalkIso (Opens.openEmbedding _) x).inv := + ((Proj.toLocallyRingedSpace 𝒜).restrictStalkIso (Opens.isOpenEmbedding _) x).inv := IsLocalization.ringHom_ext (R := A⁰_ f) ((toSpec 𝒜 f).base x).asIdeal.primeCompl (S := (Spec.structureSheaf (A⁰_ f)).presheaf.stalk ((toSpec 𝒜 f).base x)) <| (toStalk_stalkMap_toSpec _ _ _).trans <| by @@ -794,7 +795,7 @@ lemma isIso_toSpec (f) {m} (f_deg : f ∈ 𝒜 m) (hm : 0 < m) : rw [stalkMap_toSpec 𝒜 f x f_deg hm]; infer_instance haveI : LocallyRingedSpace.IsOpenImmersion (toSpec 𝒜 f) := LocallyRingedSpace.IsOpenImmersion.of_stalk_iso (toSpec 𝒜 f) - (TopCat.homeoOfIso (asIso <| (toSpec 𝒜 f).base)).openEmbedding + (TopCat.homeoOfIso (asIso <| (toSpec 𝒜 f).base)).isOpenEmbedding exact LocallyRingedSpace.IsOpenImmersion.to_iso _ end ProjectiveSpectrum.Proj diff --git a/Mathlib/AlgebraicGeometry/Properties.lean b/Mathlib/AlgebraicGeometry/Properties.lean index a5bbb08d90c15..84f113a767c02 100644 --- a/Mathlib/AlgebraicGeometry/Properties.lean +++ b/Mathlib/AlgebraicGeometry/Properties.lean @@ -39,8 +39,8 @@ instance : QuasiSober X := by quasiSober_of_open_cover (Set.range fun x => Set.range <| (X.affineCover.map x).base) · rintro ⟨_, i, rfl⟩; exact (X.affineCover.IsOpen i).base_open.isOpen_range · rintro ⟨_, i, rfl⟩ - exact @OpenEmbedding.quasiSober _ _ _ _ _ (Homeomorph.ofEmbedding _ - (X.affineCover.IsOpen i).base_open.toEmbedding).symm.openEmbedding PrimeSpectrum.quasiSober + exact @IsOpenEmbedding.quasiSober _ _ _ _ _ (Homeomorph.ofEmbedding _ + (X.affineCover.IsOpen i).base_open.toEmbedding).symm.isOpenEmbedding PrimeSpectrum.quasiSober · rw [Set.top_eq_univ, Set.sUnion_range, Set.eq_univ_iff_forall] intro x; exact ⟨_, ⟨_, rfl⟩, X.affineCover.covers x⟩ diff --git a/Mathlib/AlgebraicGeometry/Restrict.lean b/Mathlib/AlgebraicGeometry/Restrict.lean index 3eadf0265e3fc..2e8d4801e0e78 100644 --- a/Mathlib/AlgebraicGeometry/Restrict.lean +++ b/Mathlib/AlgebraicGeometry/Restrict.lean @@ -41,7 +41,7 @@ namespace Scheme.Opens /-- Open subset of a scheme as a scheme. -/ @[coe] def toScheme {X : Scheme.{u}} (U : X.Opens) : Scheme.{u} := - X.restrict U.openEmbedding + X.restrict U.isOpenEmbedding instance : CoeOut X.Opens Scheme := ⟨toScheme⟩ @@ -85,7 +85,7 @@ lemma range_ι : Set.range U.ι.base = U := Subtype.range_val lemma ι_image_top : U.ι ''ᵁ ⊤ = U := - U.openEmbedding_obj_top + U.isOpenEmbedding_obj_top lemma ι_image_le (W : U.toScheme.Opens) : U.ι ''ᵁ W ≤ U := by simp_rw [← U.ι_image_top] @@ -105,7 +105,7 @@ lemma ι_app_self : U.ι.app U = X.presheaf.map (eqToHom (X := U.ι ''ᵁ _) (by lemma eq_presheaf_map_eqToHom {V W : Opens U} (e : U.ι ''ᵁ V = U.ι ''ᵁ W) : X.presheaf.map (eqToHom e).op = - U.toScheme.presheaf.map (eqToHom <| U.openEmbedding.functor_obj_injective e).op := rfl + U.toScheme.presheaf.map (eqToHom <| U.isOpenEmbedding.functor_obj_injective e).op := rfl @[simp] lemma nonempty_iff : Nonempty U.toScheme ↔ (U : Set X).Nonempty := by @@ -121,20 +121,20 @@ def topIso : Γ(U, ⊤) ≅ Γ(X, U) := /-- The stalks of an open subscheme are isomorphic to the stalks of the original scheme. -/ def stalkIso {X : Scheme.{u}} (U : X.Opens) (x : U) : U.toScheme.presheaf.stalk x ≅ X.presheaf.stalk x.1 := - X.restrictStalkIso (Opens.openEmbedding _) _ + X.restrictStalkIso (Opens.isOpenEmbedding _) _ @[reassoc (attr := simp)] lemma germ_stalkIso_hom {X : Scheme.{u}} (U : X.Opens) {V : U.toScheme.Opens} (x : U) (hx : x ∈ V) : U.toScheme.presheaf.germ V x hx ≫ (U.stalkIso x).hom = X.presheaf.germ (U.ι ''ᵁ V) x.1 ⟨x, hx, rfl⟩ := - PresheafedSpace.restrictStalkIso_hom_eq_germ _ U.openEmbedding _ _ _ + PresheafedSpace.restrictStalkIso_hom_eq_germ _ U.isOpenEmbedding _ _ _ @[reassoc] lemma germ_stalkIso_inv {X : Scheme.{u}} (U : X.Opens) (V : U.toScheme.Opens) (x : U) (hx : x ∈ V) : X.presheaf.germ (U.ι ''ᵁ V) x ⟨x, hx, rfl⟩ ≫ (U.stalkIso x).inv = U.toScheme.presheaf.germ V x hx := - PresheafedSpace.restrictStalkIso_inv_eq_germ X.toPresheafedSpace U.openEmbedding V x hx + PresheafedSpace.restrictStalkIso_inv_eq_germ X.toPresheafedSpace U.isOpenEmbedding V x hx end Scheme.Opens @@ -168,9 +168,9 @@ instance ΓRestrictAlgebra {X : Scheme.{u}} (U : X.Opens) : lemma Scheme.map_basicOpen' (r : Γ(U, ⊤)) : U.ι ''ᵁ (U.toScheme.basicOpen r) = X.basicOpen - (X.presheaf.map (eqToHom U.openEmbedding_obj_top.symm).op r) := by - refine (Scheme.image_basicOpen (X.ofRestrict U.openEmbedding) r).trans ?_ - rw [← Scheme.basicOpen_res_eq _ _ (eqToHom U.openEmbedding_obj_top).op] + (X.presheaf.map (eqToHom U.isOpenEmbedding_obj_top.symm).op r) := by + refine (Scheme.image_basicOpen (X.ofRestrict U.isOpenEmbedding) r).trans ?_ + rw [← Scheme.basicOpen_res_eq _ _ (eqToHom U.isOpenEmbedding_obj_top).op] rw [← comp_apply, ← CategoryTheory.Functor.map_comp, ← op_comp, eqToHom_trans, eqToHom_refl, op_id, CategoryTheory.Functor.map_id] congr @@ -231,7 +231,7 @@ theorem Scheme.restrictFunctor_map_ofRestrict {U V : X.Opens} (i : U ⟶ V) : theorem Scheme.restrictFunctor_map_base {U V : X.Opens} (i : U ⟶ V) : (X.restrictFunctor.map i).left.base = (Opens.toTopCat _).map i := by ext a; refine Subtype.ext ?_ -- Porting note: `ext` did not pick up `Subtype.ext` - exact (congr_arg (fun f : X.restrict U.openEmbedding ⟶ X => f.base a) + exact (congr_arg (fun f : X.restrict U.isOpenEmbedding ⟶ X => f.base a) (X.restrictFunctor_map_ofRestrict i)) theorem Scheme.restrictFunctor_map_app_aux {U V : X.Opens} (i : U ⟶ V) (W : Opens V) : @@ -258,7 +258,7 @@ isomorphic to the structure sheaf. -/ @[simps!] def Scheme.restrictFunctorΓ : X.restrictFunctor.op ⋙ (Over.forget X).op ⋙ Scheme.Γ ≅ X.presheaf := NatIso.ofComponents - (fun U => X.presheaf.mapIso ((eqToIso (unop U).openEmbedding_obj_top).symm.op : _)) + (fun U => X.presheaf.mapIso ((eqToIso (unop U).isOpenEmbedding_obj_top).symm.op : _)) (by intro U V i dsimp @@ -348,7 +348,7 @@ theorem isPullback_morphismRestrict {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Open rw [← Category.id_comp f] refine (IsPullback.of_horiz_isIso ⟨?_⟩).paste_horiz - (IsPullback.of_hasPullback f (Y.ofRestrict U.openEmbedding)).flip + (IsPullback.of_hasPullback f (Y.ofRestrict U.isOpenEmbedding)).flip -- Porting note: changed `rw` to `erw` erw [pullbackRestrictIsoRestrict_inv_fst]; rw [Category.comp_id] @@ -434,8 +434,8 @@ theorem morphismRestrict_appLE {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) (V theorem Γ_map_morphismRestrict {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) : Scheme.Γ.map (f ∣_ U).op = - Y.presheaf.map (eqToHom U.openEmbedding_obj_top.symm).op ≫ - f.app U ≫ X.presheaf.map (eqToHom (f ⁻¹ᵁ U).openEmbedding_obj_top).op := by + Y.presheaf.map (eqToHom U.isOpenEmbedding_obj_top.symm).op ≫ + f.app U ≫ X.presheaf.map (eqToHom (f ⁻¹ᵁ U).isOpenEmbedding_obj_top).op := by rw [Scheme.Γ_map_op, morphismRestrict_app f U ⊤, f.naturality_assoc, ← X.presheaf.map_comp] rfl @@ -486,7 +486,7 @@ def morphismRestrictRestrict {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) (V : /-- Restricting a morphism twice onto a basic open set is isomorphic to one restriction. -/ def morphismRestrictRestrictBasicOpen {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) (r : Γ(Y, U)) : Arrow.mk (f ∣_ U ∣_ - U.toScheme.basicOpen (Y.presheaf.map (eqToHom U.openEmbedding_obj_top).op r)) ≅ + U.toScheme.basicOpen (Y.presheaf.map (eqToHom U.isOpenEmbedding_obj_top).op r)) ≅ Arrow.mk (f ∣_ Y.basicOpen r) := by refine morphismRestrictRestrict _ _ _ ≪≫ morphismRestrictEq _ ?_ have e := Scheme.preimage_basicOpen U.ι r diff --git a/Mathlib/AlgebraicGeometry/Scheme.lean b/Mathlib/AlgebraicGeometry/Scheme.lean index d77a37bef9383..18cc563e90561 100644 --- a/Mathlib/AlgebraicGeometry/Scheme.lean +++ b/Mathlib/AlgebraicGeometry/Scheme.lean @@ -46,7 +46,7 @@ structure Scheme extends LocallyRingedSpace where ∀ x : toLocallyRingedSpace, ∃ (U : OpenNhds x) (R : CommRingCat), Nonempty - (toLocallyRingedSpace.restrict U.openEmbedding ≅ Spec.toLocallyRingedSpace.obj (op R)) + (toLocallyRingedSpace.restrict U.isOpenEmbedding ≅ Spec.toLocallyRingedSpace.obj (op R)) namespace Scheme diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/Manifold.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/Manifold.lean index 2bc3767cffe15..45772d8d62ff6 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/Manifold.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/Manifold.lean @@ -19,10 +19,10 @@ open scoped UpperHalfPlane Manifold namespace UpperHalfPlane noncomputable instance : ChartedSpace ℂ ℍ := - UpperHalfPlane.openEmbedding_coe.singletonChartedSpace + UpperHalfPlane.isOpenEmbedding_coe.singletonChartedSpace instance : SmoothManifoldWithCorners 𝓘(ℂ) ℍ := - UpperHalfPlane.openEmbedding_coe.singleton_smoothManifoldWithCorners 𝓘(ℂ) + UpperHalfPlane.isOpenEmbedding_coe.singleton_smoothManifoldWithCorners 𝓘(ℂ) /-- The inclusion map `ℍ → ℂ` is a smooth map of manifolds. -/ theorem smooth_coe : Smooth 𝓘(ℂ) 𝓘(ℂ) ((↑) : ℍ → ℂ) := fun _ => contMDiffAt_extChartAt diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/Topology.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/Topology.lean index caecc196aeb8d..8aec80774131b 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/Topology.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/Topology.lean @@ -30,8 +30,11 @@ namespace UpperHalfPlane instance : TopologicalSpace ℍ := instTopologicalSpaceSubtype -theorem openEmbedding_coe : OpenEmbedding ((↑) : ℍ → ℂ) := - IsOpen.openEmbedding_subtype_val <| isOpen_lt continuous_const Complex.continuous_im +theorem isOpenEmbedding_coe : IsOpenEmbedding ((↑) : ℍ → ℂ) := + IsOpen.isOpenEmbedding_subtypeVal <| isOpen_lt continuous_const Complex.continuous_im + +@[deprecated (since := "2024-10-18")] +alias openEmbedding_coe := isOpenEmbedding_coe theorem embedding_coe : Embedding ((↑) : ℍ → ℂ) := embedding_subtype_val @@ -55,7 +58,7 @@ instance : T4Space ℍ := inferInstance instance : ContractibleSpace ℍ := (convex_halfspace_im_gt 0).contractibleSpace ⟨I, one_pos.trans_eq I_im.symm⟩ -instance : LocPathConnectedSpace ℍ := openEmbedding_coe.locPathConnectedSpace +instance : LocPathConnectedSpace ℍ := isOpenEmbedding_coe.locPathConnectedSpace instance : NoncompactSpace ℍ := by refine ⟨fun h => ?_⟩ @@ -65,7 +68,7 @@ instance : NoncompactSpace ℍ := by exact absurd ((this 0).1 (@left_mem_Ici ℝ _ 0)) (@lt_irrefl ℝ _ 0) instance : LocallyCompactSpace ℍ := - openEmbedding_coe.locallyCompactSpace + isOpenEmbedding_coe.locallyCompactSpace section strips @@ -118,13 +121,13 @@ theorem ModularGroup_T_zpow_mem_verticalStrip (z : ℍ) {N : ℕ} (hn : 0 < N) : end strips /-- A continuous section `ℂ → ℍ` of the natural inclusion map, bundled as a `PartialHomeomorph`. -/ -def ofComplex : PartialHomeomorph ℂ ℍ := (openEmbedding_coe.toPartialHomeomorph _).symm +def ofComplex : PartialHomeomorph ℂ ℍ := (isOpenEmbedding_coe.toPartialHomeomorph _).symm /-- Extend a function on `ℍ` arbitrarily to a function on all of `ℂ`. -/ scoped notation "↑ₕ" f => f ∘ ofComplex lemma ofComplex_apply (z : ℍ) : ofComplex (z : ℂ) = z := - OpenEmbedding.toPartialHomeomorph_left_inv .. + IsOpenEmbedding.toPartialHomeomorph_left_inv .. lemma ofComplex_apply_eq_ite (w : ℂ) : ofComplex w = if hw : 0 < w.im then ⟨w, hw⟩ else Classical.choice inferInstance := by diff --git a/Mathlib/Analysis/Normed/Ring/Units.lean b/Mathlib/Analysis/Normed/Ring/Units.lean index 86883d4c5d017..eed32209f935a 100644 --- a/Mathlib/Analysis/Normed/Ring/Units.lean +++ b/Mathlib/Analysis/Normed/Ring/Units.lean @@ -200,15 +200,18 @@ open MulOpposite Filter NormedRing /-- In a normed ring with summable geometric series, the coercion from `Rˣ` (equipped with the induced topology from the embedding in `R × R`) to `R` is an open embedding. -/ -theorem openEmbedding_val : OpenEmbedding (val : Rˣ → R) where +theorem isOpenEmbedding_val : IsOpenEmbedding (val : Rˣ → R) where toEmbedding := embedding_val_mk' (fun _ ⟨u, hu⟩ ↦ hu ▸ (inverse_continuousAt u).continuousWithinAt) Ring.inverse_unit isOpen_range := Units.isOpen +@[deprecated (since := "2024-10-18")] +alias openEmbedding_val := isOpenEmbedding_val + /-- In a normed ring with summable geometric series, the coercion from `Rˣ` (equipped with the induced topology from the embedding in `R × R`) to `R` is an open map. -/ theorem isOpenMap_val : IsOpenMap (val : Rˣ → R) := - openEmbedding_val.isOpenMap + isOpenEmbedding_val.isOpenMap end Units diff --git a/Mathlib/Analysis/SpecialFunctions/Exp.lean b/Mathlib/Analysis/SpecialFunctions/Exp.lean index 5d7dc418a5c49..9a02b4d3dc803 100644 --- a/Mathlib/Analysis/SpecialFunctions/Exp.lean +++ b/Mathlib/Analysis/SpecialFunctions/Exp.lean @@ -369,16 +369,19 @@ theorem tendsto_exp_comp_nhds_zero {f : α → ℝ} : Tendsto (fun x => exp (f x)) l (𝓝 0) ↔ Tendsto f l atBot := by simp_rw [← comp_apply (f := exp), ← tendsto_comap_iff, comap_exp_nhds_zero] -theorem openEmbedding_exp : OpenEmbedding exp := - isOpen_Ioi.openEmbedding_subtype_val.comp expOrderIso.toHomeomorph.openEmbedding +theorem isOpenEmbedding_exp : IsOpenEmbedding exp := + isOpen_Ioi.isOpenEmbedding_subtypeVal.comp expOrderIso.toHomeomorph.isOpenEmbedding + +@[deprecated (since := "2024-10-18")] +alias openEmbedding_exp := isOpenEmbedding_exp @[simp] theorem map_exp_nhds (x : ℝ) : map exp (𝓝 x) = 𝓝 (exp x) := - openEmbedding_exp.map_nhds_eq x + isOpenEmbedding_exp.map_nhds_eq x @[simp] theorem comap_exp_nhds_exp (x : ℝ) : comap exp (𝓝 (exp x)) = 𝓝 x := - (openEmbedding_exp.nhds_eq_comap x).symm + (isOpenEmbedding_exp.nhds_eq_comap x).symm theorem isLittleO_pow_exp_atTop {n : ℕ} : (fun x : ℝ => x ^ n) =o[atTop] Real.exp := by simpa [isLittleO_iff_tendsto fun x hx => ((exp_pos x).ne' hx).elim] using diff --git a/Mathlib/Geometry/Manifold/ChartedSpace.lean b/Mathlib/Geometry/Manifold/ChartedSpace.lean index 5acd6a8cea119..006514be87487 100644 --- a/Mathlib/Geometry/Manifold/ChartedSpace.lean +++ b/Mathlib/Geometry/Manifold/ChartedSpace.lean @@ -1049,7 +1049,7 @@ variable (e : PartialHomeomorph α H) /-- If a single partial homeomorphism `e` from a space `α` into `H` has source covering the whole space `α`, then that partial homeomorphism induces an `H`-charted space structure on `α`. (This condition is equivalent to `e` being an open embedding of `α` into `H`; see -`OpenEmbedding.singletonChartedSpace`.) -/ +`IsOpenEmbedding.singletonChartedSpace`.) -/ def singletonChartedSpace (h : e.source = Set.univ) : ChartedSpace H α where atlas := {e} chartAt _ := e @@ -1085,24 +1085,24 @@ theorem singleton_hasGroupoid (h : e.source = Set.univ) (G : StructureGroupoid H end PartialHomeomorph -namespace OpenEmbedding +namespace IsOpenEmbedding variable [Nonempty α] /-- An open embedding of `α` into `H` induces an `H`-charted space structure on `α`. See `PartialHomeomorph.singletonChartedSpace`. -/ -def singletonChartedSpace {f : α → H} (h : OpenEmbedding f) : ChartedSpace H α := +def singletonChartedSpace {f : α → H} (h : IsOpenEmbedding f) : ChartedSpace H α := (h.toPartialHomeomorph f).singletonChartedSpace (toPartialHomeomorph_source _ _) -theorem singletonChartedSpace_chartAt_eq {f : α → H} (h : OpenEmbedding f) {x : α} : +theorem singletonChartedSpace_chartAt_eq {f : α → H} (h : IsOpenEmbedding f) {x : α} : ⇑(@chartAt H _ α _ h.singletonChartedSpace x) = f := rfl -theorem singleton_hasGroupoid {f : α → H} (h : OpenEmbedding f) (G : StructureGroupoid H) +theorem singleton_hasGroupoid {f : α → H} (h : IsOpenEmbedding f) (G : StructureGroupoid H) [ClosedUnderRestriction G] : @HasGroupoid _ _ _ _ h.singletonChartedSpace G := (h.toPartialHomeomorph f).singleton_hasGroupoid (toPartialHomeomorph_source _ _) G -end OpenEmbedding +end IsOpenEmbedding end Singleton diff --git a/Mathlib/Geometry/Manifold/ContMDiff/Basic.lean b/Mathlib/Geometry/Manifold/ContMDiff/Basic.lean index cd4f4365ba286..b79ff42cc1ad9 100644 --- a/Mathlib/Geometry/Manifold/ContMDiff/Basic.lean +++ b/Mathlib/Geometry/Manifold/ContMDiff/Basic.lean @@ -13,7 +13,7 @@ In this file, we show that standard operations on smooth maps between smooth man * `contMDiff_id` gives the smoothness of the identity * `contMDiff_const` gives the smoothness of constant functions * `contMDiff_inclusion` shows that the inclusion between open sets of a topological space is smooth -* `contMDiff_openEmbedding` shows that if `M` has a `ChartedSpace` structure induced by an open +* `contMDiff_isOpenEmbedding` shows that if `M` has a `ChartedSpace` structure induced by an open embedding `e : M → H`, then `e` is smooth. ## Tags @@ -358,11 +358,11 @@ end ChartedSpace section -variable {e : M → H} (h : OpenEmbedding e) {n : WithTop ℕ} +variable {e : M → H} (h : IsOpenEmbedding e) {n : WithTop ℕ} /-- If the `ChartedSpace` structure on a manifold `M` is given by an open embedding `e : M → H`, then `e` is smooth. -/ -lemma contMDiff_openEmbedding [Nonempty M] : +lemma contMDiff_isOpenEmbedding [Nonempty M] : haveI := h.singletonChartedSpace; ContMDiff I I n e := by haveI := h.singleton_smoothManifoldWithCorners I rw [@contMDiff_iff _ _ _ _ _ _ _ _ _ _ h.singletonChartedSpace] @@ -385,12 +385,15 @@ lemma contMDiff_openEmbedding [Nonempty M] : h.toPartialHomeomorph_target] at this exact this +@[deprecated (since := "2024-10-18")] +alias contMDiff_openEmbedding := contMDiff_isOpenEmbedding + variable {I I'} /-- If the `ChartedSpace` structure on a manifold `M` is given by an open embedding `e : M → H`, then the inverse of `e` is smooth. -/ -lemma contMDiffOn_openEmbedding_symm [Nonempty M] : +lemma contMDiffOn_isOpenEmbedding_symm [Nonempty M] : haveI := h.singletonChartedSpace; ContMDiffOn I I - n (OpenEmbedding.toPartialHomeomorph e h).symm (range e) := by + n (IsOpenEmbedding.toPartialHomeomorph e h).symm (range e) := by haveI := h.singleton_smoothManifoldWithCorners I rw [@contMDiffOn_iff] constructor @@ -409,21 +412,27 @@ lemma contMDiffOn_openEmbedding_symm [Nonempty M] : apply I.right_inv exact mem_of_subset_of_mem (extChartAt_target_subset_range _ _) hz.1 +@[deprecated (since := "2024-10-18")] +alias contMDiffOn_openEmbedding_symm := contMDiffOn_isOpenEmbedding_symm + variable [ChartedSpace H M] -variable [Nonempty M'] {e' : M' → H'} (h' : OpenEmbedding e') +variable [Nonempty M'] {e' : M' → H'} (h' : IsOpenEmbedding e') /-- Let `M'` be a manifold whose chart structure is given by an open embedding `e'` into its model space `H'`. Then the smoothness of `e' ∘ f : M → H'` implies the smoothness of `f`. This is useful, for example, when `e' ∘ f = g ∘ e` for smooth maps `e : M → X` and `g : X → H'`. -/ -lemma ContMDiff.of_comp_openEmbedding {f : M → M'} (hf : ContMDiff I I' n (e' ∘ f)) : +lemma ContMDiff.of_comp_isOpenEmbedding {f : M → M'} (hf : ContMDiff I I' n (e' ∘ f)) : haveI := h'.singletonChartedSpace; ContMDiff I I' n f := by have : f = (h'.toPartialHomeomorph e').symm ∘ e' ∘ f := by ext - rw [Function.comp_apply, Function.comp_apply, OpenEmbedding.toPartialHomeomorph_left_inv] + rw [Function.comp_apply, Function.comp_apply, IsOpenEmbedding.toPartialHomeomorph_left_inv] rw [this] apply @ContMDiffOn.comp_contMDiff _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ - h'.singletonChartedSpace _ _ (range e') _ (contMDiffOn_openEmbedding_symm h') hf + h'.singletonChartedSpace _ _ (range e') _ (contMDiffOn_isOpenEmbedding_symm h') hf simp +@[deprecated (since := "2024-10-18")] +alias ContMDiff.of_comp_openEmbedding := ContMDiff.of_comp_isOpenEmbedding + end diff --git a/Mathlib/Geometry/Manifold/Instances/UnitsOfNormedAlgebra.lean b/Mathlib/Geometry/Manifold/Instances/UnitsOfNormedAlgebra.lean index e7690b95e11f0..afe2f7f2112cb 100644 --- a/Mathlib/Geometry/Manifold/Instances/UnitsOfNormedAlgebra.lean +++ b/Mathlib/Geometry/Manifold/Instances/UnitsOfNormedAlgebra.lean @@ -33,7 +33,7 @@ namespace Units variable {R : Type*} [NormedRing R] [CompleteSpace R] instance : ChartedSpace R Rˣ := - openEmbedding_val.singletonChartedSpace + isOpenEmbedding_val.singletonChartedSpace theorem chartAt_apply {a : Rˣ} {b : Rˣ} : chartAt R a b = b := rfl @@ -44,17 +44,17 @@ theorem chartAt_source {a : Rˣ} : (chartAt R a).source = Set.univ := variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] [NormedAlgebra 𝕜 R] instance : SmoothManifoldWithCorners 𝓘(𝕜, R) Rˣ := - openEmbedding_val.singleton_smoothManifoldWithCorners 𝓘(𝕜, R) + isOpenEmbedding_val.singleton_smoothManifoldWithCorners 𝓘(𝕜, R) /-- For a complete normed ring `R`, the embedding of the units `Rˣ` into `R` is a smooth map between manifolds. -/ lemma contMDiff_val {m : ℕ∞} : ContMDiff 𝓘(𝕜, R) 𝓘(𝕜, R) m (val : Rˣ → R) := - contMDiff_openEmbedding 𝓘(𝕜, R) Units.openEmbedding_val + contMDiff_isOpenEmbedding 𝓘(𝕜, R) Units.isOpenEmbedding_val /-- The units of a complete normed ring form a Lie group. -/ instance : LieGroup 𝓘(𝕜, R) Rˣ where smooth_mul := by - apply ContMDiff.of_comp_openEmbedding Units.openEmbedding_val + apply ContMDiff.of_comp_isOpenEmbedding Units.isOpenEmbedding_val have : (val : Rˣ → R) ∘ (fun x : Rˣ × Rˣ => x.1 * x.2) = (fun x : R × R => x.1 * x.2) ∘ (fun x : Rˣ × Rˣ => (x.1, x.2)) := by ext; simp rw [this] @@ -65,7 +65,7 @@ instance : LieGroup 𝓘(𝕜, R) Rˣ where rw [contMDiff_iff_contDiff] exact contDiff_mul smooth_inv := by - apply ContMDiff.of_comp_openEmbedding Units.openEmbedding_val + apply ContMDiff.of_comp_isOpenEmbedding Units.isOpenEmbedding_val have : (val : Rˣ → R) ∘ (fun x : Rˣ => x⁻¹) = Ring.inverse ∘ val := by ext; simp rw [this, ContMDiff] refine fun x => ContMDiffAt.comp x ?_ (contMDiff_val x) diff --git a/Mathlib/Geometry/Manifold/LocalInvariantProperties.lean b/Mathlib/Geometry/Manifold/LocalInvariantProperties.lean index 5ea4b4a42e6b8..6b89039b66f57 100644 --- a/Mathlib/Geometry/Manifold/LocalInvariantProperties.lean +++ b/Mathlib/Geometry/Manifold/LocalInvariantProperties.lean @@ -490,7 +490,7 @@ theorem liftPropAt_iff_comp_subtype_val (hG : LocalInvariantProp G G' P) {U : Op LiftPropAt P f x ↔ LiftPropAt P (f ∘ Subtype.val) x := by simp only [LiftPropAt, liftPropWithinAt_iff'] congrm ?_ ∧ ?_ - · simp_rw [continuousWithinAt_univ, U.openEmbedding'.continuousAt_iff] + · simp_rw [continuousWithinAt_univ, U.isOpenEmbedding'.continuousAt_iff] · apply hG.congr_iff exact (U.chartAt_subtype_val_symm_eventuallyEq).fun_comp (chartAt H' (f x) ∘ f) @@ -500,7 +500,7 @@ theorem liftPropAt_iff_comp_inclusion (hG : LocalInvariantProp G G' P) {U V : Op simp only [LiftPropAt, liftPropWithinAt_iff'] congrm ?_ ∧ ?_ · simp_rw [continuousWithinAt_univ, - (TopologicalSpace.Opens.openEmbedding_of_le hUV).continuousAt_iff] + (TopologicalSpace.Opens.isOpenEmbedding_of_le hUV).continuousAt_iff] · apply hG.congr_iff exact (TopologicalSpace.Opens.chartAt_inclusion_symm_eventuallyEq hUV).fun_comp (chartAt H' (f (Set.inclusion hUV x)) ∘ f) diff --git a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean index d6cebfb34829f..ac7279ed0be9c 100644 --- a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean +++ b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean @@ -724,13 +724,17 @@ theorem PartialHomeomorph.singleton_smoothManifoldWithCorners @SmoothManifoldWithCorners.mk' _ _ _ _ _ _ _ _ _ _ (id _) <| e.singleton_hasGroupoid h (contDiffGroupoid ∞ I) -theorem OpenEmbedding.singleton_smoothManifoldWithCorners {𝕜 : Type*} [NontriviallyNormedField 𝕜] +theorem IsOpenEmbedding.singleton_smoothManifoldWithCorners {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] (I : ModelWithCorners 𝕜 E H) {M : Type*} [TopologicalSpace M] [Nonempty M] {f : M → H} - (h : OpenEmbedding f) : + (h : IsOpenEmbedding f) : @SmoothManifoldWithCorners 𝕜 _ E _ _ H _ I M _ h.singletonChartedSpace := (h.toPartialHomeomorph f).singleton_smoothManifoldWithCorners I (by simp) +@[deprecated (since := "2024-10-18")] +alias OpenEmbedding.singleton_smoothManifoldWithCorners := + IsOpenEmbedding.singleton_smoothManifoldWithCorners + namespace TopologicalSpace.Opens open TopologicalSpace diff --git a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace.lean b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace.lean index f0aedb017be00..2fc0a8b8893ef 100644 --- a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace.lean +++ b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace.lean @@ -217,8 +217,8 @@ instance is_sheafedSpace_iso {X Y : LocallyRingedSpace.{u}} (f : X ⟶ Y) [IsIso /-- The restriction of a locally ringed space along an open embedding. -/ @[simps!] -def restrict {U : TopCat} (X : LocallyRingedSpace.{u}) {f : U ⟶ X.toTopCat} (h : OpenEmbedding f) : - LocallyRingedSpace where +def restrict {U : TopCat} (X : LocallyRingedSpace.{u}) {f : U ⟶ X.toTopCat} + (h : IsOpenEmbedding f) : LocallyRingedSpace where localRing := by intro x -- We show that the stalk of the restriction is isomorphic to the original stalk, @@ -228,13 +228,13 @@ def restrict {U : TopCat} (X : LocallyRingedSpace.{u}) {f : U ⟶ X.toTopCat} (h /-- The canonical map from the restriction to the subspace. -/ def ofRestrict {U : TopCat} (X : LocallyRingedSpace.{u}) - {f : U ⟶ X.toTopCat} (h : OpenEmbedding f) : X.restrict h ⟶ X := + {f : U ⟶ X.toTopCat} (h : IsOpenEmbedding f) : X.restrict h ⟶ X := ⟨X.toPresheafedSpace.ofRestrict h, fun _ => inferInstance⟩ /-- The restriction of a locally ringed space `X` to the top subspace is isomorphic to `X` itself. -/ def restrictTopIso (X : LocallyRingedSpace.{u}) : - X.restrict (Opens.openEmbedding ⊤) ≅ X := + X.restrict (Opens.isOpenEmbedding ⊤) ≅ X := isoOfSheafedSpaceIso X.toSheafedSpace.restrictTopIso /-- The global sections, notated Gamma. @@ -436,7 +436,7 @@ theorem preimage_basicOpen {X Y : LocallyRingedSpace.{u}} (f : X ⟶ Y) {U : Ope rw [← stalkMap_germ_apply] at hx exact (isUnit_map_iff (f.toShHom.stalkMap _) _).mp hx -variable {U : TopCat} (X : LocallyRingedSpace.{u}) {f : U ⟶ X.toTopCat} (h : OpenEmbedding f) +variable {U : TopCat} (X : LocallyRingedSpace.{u}) {f : U ⟶ X.toTopCat} (h : IsOpenEmbedding f) (V : Opens U) (x : U) (hx : x ∈ V) /-- For an open embedding `f : U ⟶ X` and a point `x : U`, we get an isomorphism between the stalk diff --git a/Mathlib/Geometry/RingedSpace/OpenImmersion.lean b/Mathlib/Geometry/RingedSpace/OpenImmersion.lean index 11dbae9b63d30..f549a8e9ee841 100644 --- a/Mathlib/Geometry/RingedSpace/OpenImmersion.lean +++ b/Mathlib/Geometry/RingedSpace/OpenImmersion.lean @@ -68,7 +68,7 @@ spaces, such that the sheaf map `Y(V) ⟶ f _* X(V)` is an iso for each `V ⊆ U class PresheafedSpace.IsOpenImmersion {X Y : PresheafedSpace C} (f : X ⟶ Y) : Prop where /-- the underlying continuous map of underlying spaces from the source to an open subset of the target. -/ - base_open : OpenEmbedding f.base + base_open : IsOpenEmbedding f.base /-- the underlying sheaf morphism is an isomorphism on each open subset -/ c_iso : ∀ U : Opens X, IsIso (f.c.app (op (base_open.isOpenMap.functor.obj U))) @@ -214,7 +214,7 @@ theorem app_inv_app' (U : Opens Y) (hU : (U : Set Y) ⊆ Set.range f.base) : /-- An isomorphism is an open immersion. -/ instance ofIso {X Y : PresheafedSpace C} (H : X ≅ Y) : IsOpenImmersion H.hom where - base_open := (TopCat.homeoOfIso ((forget C).mapIso H)).openEmbedding + base_open := (TopCat.homeoOfIso ((forget C).mapIso H)).isOpenEmbedding -- Porting note: `inferInstance` will fail if Lean is not told that `H.hom.c` is iso c_iso _ := letI : IsIso H.hom.c := c_isIso_of_iso H.hom; inferInstance @@ -223,7 +223,7 @@ instance (priority := 100) ofIsIso {X Y : PresheafedSpace C} (f : X ⟶ Y) [IsIs AlgebraicGeometry.PresheafedSpace.IsOpenImmersion.ofIso (asIso f) instance ofRestrict {X : TopCat} (Y : PresheafedSpace C) {f : X ⟶ Y.carrier} - (hf : OpenEmbedding f) : IsOpenImmersion (Y.ofRestrict hf) where + (hf : IsOpenEmbedding f) : IsOpenImmersion (Y.ofRestrict hf) where base_open := hf c_iso U := by dsimp @@ -243,7 +243,7 @@ instance ofRestrict {X : TopCat} (Y : PresheafedSpace C) {f : X ⟶ Y.carrier} @[elementwise, simp] theorem ofRestrict_invApp {C : Type*} [Category C] (X : PresheafedSpace C) {Y : TopCat} - {f : Y ⟶ TopCat.of X.carrier} (h : OpenEmbedding f) (U : Opens (X.restrict h).carrier) : + {f : Y ⟶ TopCat.of X.carrier} (h : IsOpenEmbedding f) (U : Opens (X.restrict h).carrier) : (PresheafedSpace.IsOpenImmersion.ofRestrict X h).invApp _ U = 𝟙 _ := by delta invApp rw [IsIso.comp_inv_eq, Category.id_comp] @@ -286,7 +286,7 @@ variable {X Y Z : PresheafedSpace C} (f : X ⟶ Z) [hf : IsOpenImmersion f] (g : /-- (Implementation.) The projection map when constructing the pullback along an open immersion. -/ def pullbackConeOfLeftFst : - Y.restrict (TopCat.snd_openEmbedding_of_left_openEmbedding hf.base_open g.base) ⟶ X where + Y.restrict (TopCat.snd_isOpenEmbedding_of_left hf.base_open g.base) ⟶ X where base := pullback.fst _ _ c := { app := fun U => @@ -377,7 +377,7 @@ def pullbackConeOfLeftLift : s.pt ⟶ (pullbackConeOfLeft f g).pt where dsimp [s'] rw [Function.comp_def, ← Set.preimage_preimage] rw [Set.preimage_image_eq _ - (TopCat.snd_openEmbedding_of_left_openEmbedding hf.base_open g.base).inj] + (TopCat.snd_isOpenEmbedding_of_left hf.base_open g.base).inj] rfl)) naturality := fun U V i => by erw [s.snd.c.naturality_assoc] @@ -727,7 +727,7 @@ whose forgetful functor reflects isomorphisms, preserves limits and filtered col Then a morphism `X ⟶ Y` that is a topological open embedding is an open immersion iff every stalk map is an iso. -/ -theorem of_stalk_iso {X Y : SheafedSpace C} (f : X ⟶ Y) (hf : OpenEmbedding f.base) +theorem of_stalk_iso {X Y : SheafedSpace C} (f : X ⟶ Y) (hf : IsOpenEmbedding f.base) [H : ∀ x : X.1, IsIso (f.stalkMap x)] : SheafedSpace.IsOpenImmersion f := { base_open := hf c_iso := fun U => by @@ -738,7 +738,7 @@ theorem of_stalk_iso {X Y : SheafedSpace C} (f : X ⟶ Y) (hf : OpenEmbedding f. specialize H y delta PresheafedSpace.Hom.stalkMap at H haveI H' := - TopCat.Presheaf.stalkPushforward.stalkPushforward_iso_of_openEmbedding C hf X.presheaf y + TopCat.Presheaf.stalkPushforward.stalkPushforward_iso_of_isOpenEmbedding C hf X.presheaf y have := IsIso.comp_isIso' H (@IsIso.inv_isIso _ _ _ _ _ H') rwa [Category.assoc, IsIso.hom_inv_id, Category.comp_id] at this } @@ -812,12 +812,12 @@ theorem app_inv_app' (U : Opens Y) (hU : (U : Set Y) ⊆ Set.range f.base) : PresheafedSpace.IsOpenImmersion.app_invApp f U instance ofRestrict {X : TopCat} (Y : SheafedSpace C) {f : X ⟶ Y.carrier} - (hf : OpenEmbedding f) : IsOpenImmersion (Y.ofRestrict hf) := + (hf : IsOpenEmbedding f) : IsOpenImmersion (Y.ofRestrict hf) := PresheafedSpace.IsOpenImmersion.ofRestrict _ hf @[elementwise, simp] theorem ofRestrict_invApp {C : Type*} [Category C] (X : SheafedSpace C) {Y : TopCat} - {f : Y ⟶ TopCat.of X.carrier} (h : OpenEmbedding f) (U : Opens (X.restrict h).carrier) : + {f : Y ⟶ TopCat.of X.carrier} (h : IsOpenEmbedding f) (U : Opens (X.restrict h).carrier) : (SheafedSpace.IsOpenImmersion.ofRestrict X h).invApp _ U = 𝟙 _ := PresheafedSpace.IsOpenImmersion.ofRestrict_invApp _ h U @@ -839,7 +839,7 @@ section Prod variable [HasLimits C] {ι : Type v} (F : Discrete ι ⥤ SheafedSpace.{_, v, v} C) [HasColimit F] (i : Discrete ι) -theorem sigma_ι_openEmbedding : OpenEmbedding (colimit.ι F i).base := by +theorem sigma_ι_isOpenEmbedding : IsOpenEmbedding (colimit.ι F i).base := by rw [← show _ = (colimit.ι F i).base from ι_preservesColimitsIso_inv (SheafedSpace.forget C) F i] have : _ = _ ≫ colimit.ι (Discrete.functor ((F ⋙ SheafedSpace.forget C).obj ∘ Discrete.mk)) i := HasColimit.isoOfNatIso_ι_hom Discrete.natIsoFunctor i @@ -850,17 +850,20 @@ theorem sigma_ι_openEmbedding : OpenEmbedding (colimit.ι F i).base := by rw [← Iso.eq_comp_inv] at this cases i rw [this, ← Category.assoc] - -- Porting note: `simp_rw` can't use `TopCat.openEmbedding_iff_comp_isIso` and - -- `TopCat.openEmbedding_iff_isIso_comp`. + -- Porting note: `simp_rw` can't use `TopCat.isOpenEmbedding_iff_comp_isIso` and + -- `TopCat.isOpenEmbedding_iff_isIso_comp`. -- See https://github.com/leanprover-community/mathlib4/issues/5026 - erw [TopCat.openEmbedding_iff_comp_isIso, TopCat.openEmbedding_iff_comp_isIso, - TopCat.openEmbedding_iff_comp_isIso, TopCat.openEmbedding_iff_isIso_comp] - exact openEmbedding_sigmaMk + erw [TopCat.isOpenEmbedding_iff_comp_isIso, TopCat.isOpenEmbedding_iff_comp_isIso, + TopCat.isOpenEmbedding_iff_comp_isIso, TopCat.isOpenEmbedding_iff_isIso_comp] + exact isOpenEmbedding_sigmaMk + +@[deprecated (since := "2024-10-18")] +alias sigma_ι_openEmbedding := sigma_ι_isOpenEmbedding theorem image_preimage_is_empty (j : Discrete ι) (h : i ≠ j) (U : Opens (F.obj i)) : (Opens.map (colimit.ι (F ⋙ SheafedSpace.forgetToPresheafedSpace) j).base).obj ((Opens.map (preservesColimitIso SheafedSpace.forgetToPresheafedSpace F).inv.base).obj - ((sigma_ι_openEmbedding F i).isOpenMap.functor.obj U)) = + ((sigma_ι_isOpenEmbedding F i).isOpenMap.functor.obj U)) = ⊥ := by ext x apply iff_false_intro @@ -880,15 +883,15 @@ theorem image_preimage_is_empty (j : Discrete ι) (h : i ≠ j) (U : Opens (F.ob instance sigma_ι_isOpenImmersion [HasStrictTerminalObjects C] : SheafedSpace.IsOpenImmersion (colimit.ι F i) where - base_open := sigma_ι_openEmbedding F i + base_open := sigma_ι_isOpenEmbedding F i c_iso U := by have e : colimit.ι F i = _ := (ι_preservesColimitsIso_inv SheafedSpace.forgetToPresheafedSpace F i).symm have H : - OpenEmbedding + IsOpenEmbedding (colimit.ι (F ⋙ SheafedSpace.forgetToPresheafedSpace) i ≫ (preservesColimitIso SheafedSpace.forgetToPresheafedSpace F).inv).base := - e ▸ sigma_ι_openEmbedding F i + e ▸ sigma_ι_isOpenEmbedding F i suffices IsIso <| (colimit.ι (F ⋙ SheafedSpace.forgetToPresheafedSpace) i ≫ (preservesColimitIso SheafedSpace.forgetToPresheafedSpace F).inv).c.app <| op (H.isOpenMap.functor.obj U) by @@ -921,7 +924,7 @@ end SheafedSpace.IsOpenImmersion namespace LocallyRingedSpace.IsOpenImmersion -instance (X : LocallyRingedSpace) {U : TopCat} (f : U ⟶ X.toTopCat) (hf : OpenEmbedding f) : +instance (X : LocallyRingedSpace) {U : TopCat} (f : U ⟶ X.toTopCat) (hf : IsOpenEmbedding f) : LocallyRingedSpace.IsOpenImmersion (X.ofRestrict hf) := PresheafedSpace.IsOpenImmersion.ofRestrict X.toPresheafedSpace hf @@ -951,7 +954,7 @@ set_option synthInstance.maxHeartbeats 40000 in /-- An explicit pullback cone over `cospan f g` if `f` is an open immersion. -/ def pullbackConeOfLeft : PullbackCone f g := by refine PullbackCone.mk ?_ - (Y.ofRestrict (TopCat.snd_openEmbedding_of_left_openEmbedding H.base_open g.base)) ?_ + (Y.ofRestrict (TopCat.snd_isOpenEmbedding_of_left H.base_open g.base)) ?_ · use PresheafedSpace.IsOpenImmersion.pullbackConeOfLeftFst f.1 g.1 intro x have := PresheafedSpace.stalkMap.congr_hom _ _ @@ -1154,7 +1157,7 @@ whose forgetful functor reflects isomorphisms, preserves limits and filtered col Then a morphism `X ⟶ Y` that is a topological open embedding is an open immersion iff every stalk map is an iso. -/ -theorem of_stalk_iso {X Y : LocallyRingedSpace} (f : X ⟶ Y) (hf : OpenEmbedding f.base) +theorem of_stalk_iso {X Y : LocallyRingedSpace} (f : X ⟶ Y) (hf : IsOpenEmbedding f.base) [stalk_iso : ∀ x : X.1, IsIso (f.stalkMap x)] : LocallyRingedSpace.IsOpenImmersion f := SheafedSpace.IsOpenImmersion.of_stalk_iso _ hf (H := stalk_iso) @@ -1223,12 +1226,12 @@ theorem app_inv_app' (U : Opens Y) (hU : (U : Set Y) ⊆ Set.range f.base) : PresheafedSpace.IsOpenImmersion.app_invApp f.1 U instance ofRestrict {X : TopCat} (Y : LocallyRingedSpace) {f : X ⟶ Y.carrier} - (hf : OpenEmbedding f) : IsOpenImmersion (Y.ofRestrict hf) := + (hf : IsOpenEmbedding f) : IsOpenImmersion (Y.ofRestrict hf) := PresheafedSpace.IsOpenImmersion.ofRestrict _ hf @[elementwise, simp] theorem ofRestrict_invApp (X : LocallyRingedSpace) {Y : TopCat} - {f : Y ⟶ TopCat.of X.carrier} (h : OpenEmbedding f) (U : Opens (X.restrict h).carrier) : + {f : Y ⟶ TopCat.of X.carrier} (h : IsOpenEmbedding f) (U : Opens (X.restrict h).carrier) : (LocallyRingedSpace.IsOpenImmersion.ofRestrict X h).invApp _ U = 𝟙 _ := PresheafedSpace.IsOpenImmersion.ofRestrict_invApp _ h U diff --git a/Mathlib/Geometry/RingedSpace/PresheafedSpace.lean b/Mathlib/Geometry/RingedSpace/PresheafedSpace.lean index 12c2d277fbd66..366aeeda7b4b0 100644 --- a/Mathlib/Geometry/RingedSpace/PresheafedSpace.lean +++ b/Mathlib/Geometry/RingedSpace/PresheafedSpace.lean @@ -293,7 +293,7 @@ section Restrict -/ @[simps] def restrict {U : TopCat} (X : PresheafedSpace C) {f : U ⟶ (X : TopCat)} - (h : OpenEmbedding f) : PresheafedSpace C where + (h : IsOpenEmbedding f) : PresheafedSpace C where carrier := U presheaf := h.isOpenMap.functor.op ⋙ X.presheaf @@ -301,7 +301,7 @@ def restrict {U : TopCat} (X : PresheafedSpace C) {f : U ⟶ (X : TopCat)} -/ @[simps] def ofRestrict {U : TopCat} (X : PresheafedSpace C) {f : U ⟶ (X : TopCat)} - (h : OpenEmbedding f) : X.restrict h ⟶ X where + (h : IsOpenEmbedding f) : X.restrict h ⟶ X where base := f c := { app := fun V => X.presheaf.map (h.isOpenMap.adjunction.counit.app V.unop).op @@ -310,8 +310,8 @@ def ofRestrict {U : TopCat} (X : PresheafedSpace C) {f : U ⟶ (X : TopCat)} rw [← map_comp, ← map_comp] rfl } -instance ofRestrict_mono {U : TopCat} (X : PresheafedSpace C) (f : U ⟶ X.1) (hf : OpenEmbedding f) : - Mono (X.ofRestrict hf) := by +instance ofRestrict_mono {U : TopCat} (X : PresheafedSpace C) (f : U ⟶ X.1) + (hf : IsOpenEmbedding f) : Mono (X.ofRestrict hf) := by haveI : Mono f := (TopCat.mono_iff_injective _).mpr hf.inj constructor intro Z g₁ g₂ eq @@ -339,14 +339,14 @@ instance ofRestrict_mono {U : TopCat} (X : PresheafedSpace C) (f : U ⟶ X.1) (h simpa using h theorem restrict_top_presheaf (X : PresheafedSpace C) : - (X.restrict (Opens.openEmbedding ⊤)).presheaf = + (X.restrict (Opens.isOpenEmbedding ⊤)).presheaf = (Opens.inclusionTopIso X.carrier).inv _* X.presheaf := by dsimp rw [Opens.inclusion'_top_functor X.carrier] rfl theorem ofRestrict_top_c (X : PresheafedSpace C) : - (X.ofRestrict (Opens.openEmbedding ⊤)).c = + (X.ofRestrict (Opens.isOpenEmbedding ⊤)).c = eqToHom (by rw [restrict_top_presheaf, ← Presheaf.Pushforward.comp_eq] @@ -365,14 +365,14 @@ theorem ofRestrict_top_c (X : PresheafedSpace C) : subspace. -/ @[simps] -def toRestrictTop (X : PresheafedSpace C) : X ⟶ X.restrict (Opens.openEmbedding ⊤) where +def toRestrictTop (X : PresheafedSpace C) : X ⟶ X.restrict (Opens.isOpenEmbedding ⊤) where base := (Opens.inclusionTopIso X.carrier).inv c := eqToHom (restrict_top_presheaf X) /-- The isomorphism from the restriction to the top subspace. -/ @[simps] -def restrictTopIso (X : PresheafedSpace C) : X.restrict (Opens.openEmbedding ⊤) ≅ X where +def restrictTopIso (X : PresheafedSpace C) : X.restrict (Opens.isOpenEmbedding ⊤) ≅ X where hom := X.ofRestrict _ inv := X.toRestrictTop hom_inv_id := by diff --git a/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean b/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean index bcfd0917ee844..16c8a58289001 100644 --- a/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean +++ b/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean @@ -121,14 +121,17 @@ abbrev toTopGlueData : TopCat.GlueData := { f_open := fun i j => (D.f_open i j).base_open toGlueData := 𝖣.mapGlueData (forget C) } -theorem ι_openEmbedding [HasLimits C] (i : D.J) : OpenEmbedding (𝖣.ι i).base := by +theorem ι_isOpenEmbedding [HasLimits C] (i : D.J) : IsOpenEmbedding (𝖣.ι i).base := by rw [← show _ = (𝖣.ι i).base from 𝖣.ι_gluedIso_inv (PresheafedSpace.forget _) _] -- Porting note: added this erewrite erw [coe_comp] refine - OpenEmbedding.comp - (TopCat.homeoOfIso (𝖣.gluedIso (PresheafedSpace.forget _)).symm).openEmbedding - (D.toTopGlueData.ι_openEmbedding i) + IsOpenEmbedding.comp + (TopCat.homeoOfIso (𝖣.gluedIso (PresheafedSpace.forget _)).symm).isOpenEmbedding + (D.toTopGlueData.ι_isOpenEmbedding i) + +@[deprecated (since := "2024-10-18")] +alias ι_openEmbedding := ι_isOpenEmbedding theorem pullback_base (i j k : D.J) (S : Set (D.V (i, j)).carrier) : (π₂ i, j, k) '' ((π₁ i, j, k) ⁻¹' S) = D.f i k ⁻¹' (D.f i j '' S) := by @@ -237,7 +240,7 @@ theorem snd_invApp_t_app (i j k : D.J) (U : Opens (pullback (D.f i j) (D.f i k)) variable [HasLimits C] theorem ι_image_preimage_eq (i j : D.J) (U : Opens (D.U i).carrier) : - (Opens.map (𝖣.ι j).base).obj ((D.ι_openEmbedding i).isOpenMap.functor.obj U) = + (Opens.map (𝖣.ι j).base).obj ((D.ι_isOpenEmbedding i).isOpenMap.functor.obj U) = (opensFunctor (D.f j i)).obj ((Opens.map (𝖣.t j i).base).obj ((Opens.map (𝖣.f i j).base).obj U)) := by ext1 @@ -265,7 +268,7 @@ theorem ι_image_preimage_eq (i j : D.J) (U : Opens (D.U i).carrier) : def opensImagePreimageMap (i j : D.J) (U : Opens (D.U i).carrier) : (D.U i).presheaf.obj (op U) ⟶ (D.U j).presheaf.obj (op <| - (Opens.map (𝖣.ι j).base).obj ((D.ι_openEmbedding i).isOpenMap.functor.obj U)) := + (Opens.map (𝖣.ι j).base).obj ((D.ι_isOpenEmbedding i).isOpenMap.functor.obj U)) := (D.f i j).c.app (op U) ≫ (D.t j i).c.app _ ≫ (D.f_open j i).invApp _ (unop _) ≫ @@ -311,7 +314,7 @@ the image `ι '' U` in the glued space is the limit of this diagram. -/ abbrev diagramOverOpen {i : D.J} (U : Opens (D.U i).carrier) : -- Porting note : ↓ these need to be explicit (WalkingMultispan D.diagram.fstFrom D.diagram.sndFrom)ᵒᵖ ⥤ C := - componentwiseDiagram 𝖣.diagram.multispan ((D.ι_openEmbedding i).isOpenMap.functor.obj U) + componentwiseDiagram 𝖣.diagram.multispan ((D.ι_isOpenEmbedding i).isOpenMap.functor.obj U) /-- (Implementation) The projection from the limit of `diagram_over_open` to a component of `D.U j`. -/ @@ -490,7 +493,7 @@ instance componentwise_diagram_π_isIso (i : D.J) (U : Opens (D.U i).carrier) : exact Iso.inv_hom_id ((D.U i).presheaf.mapIso (eqToIso _)) instance ιIsOpenImmersion (i : D.J) : IsOpenImmersion (𝖣.ι i) where - base_open := D.ι_openEmbedding i + base_open := D.ι_isOpenEmbedding i c_iso U := by erw [← colimitPresheafObjIsoComponentwiseLimit_hom_π]; infer_instance /-- The following diagram is a pullback, i.e. `Vᵢⱼ` is the intersection of `Uᵢ` and `Uⱼ` in `X`. diff --git a/Mathlib/Geometry/RingedSpace/SheafedSpace.lean b/Mathlib/Geometry/RingedSpace/SheafedSpace.lean index 3af5f2d7e91d0..4a94f1819e128 100644 --- a/Mathlib/Geometry/RingedSpace/SheafedSpace.lean +++ b/Mathlib/Geometry/RingedSpace/SheafedSpace.lean @@ -156,20 +156,20 @@ open TopCat.Presheaf /-- The restriction of a sheafed space along an open embedding into the space. -/ -def restrict {U : TopCat} (X : SheafedSpace C) {f : U ⟶ (X : TopCat)} (h : OpenEmbedding f) : +def restrict {U : TopCat} (X : SheafedSpace C) {f : U ⟶ (X : TopCat)} (h : IsOpenEmbedding f) : SheafedSpace C := - { X.toPresheafedSpace.restrict h with IsSheaf := isSheaf_of_openEmbedding h X.IsSheaf } + { X.toPresheafedSpace.restrict h with IsSheaf := isSheaf_of_isOpenEmbedding h X.IsSheaf } /-- The map from the restriction of a presheafed space. -/ @[simps!] def ofRestrict {U : TopCat} (X : SheafedSpace C) {f : U ⟶ (X : TopCat)} - (h : OpenEmbedding f) : X.restrict h ⟶ X := X.toPresheafedSpace.ofRestrict h + (h : IsOpenEmbedding f) : X.restrict h ⟶ X := X.toPresheafedSpace.ofRestrict h /-- The restriction of a sheafed space `X` to the top subspace is isomorphic to `X` itself. -/ @[simps! hom inv] -def restrictTopIso (X : SheafedSpace C) : X.restrict (Opens.openEmbedding ⊤) ≅ X := +def restrictTopIso (X : SheafedSpace C) : X.restrict (Opens.isOpenEmbedding ⊤) ≅ X := isoMk (X.toPresheafedSpace.restrictTopIso) /-- The global sections, notated Gamma. diff --git a/Mathlib/Geometry/RingedSpace/Stalks.lean b/Mathlib/Geometry/RingedSpace/Stalks.lean index 12c20e0348392..8f6f8f792d874 100644 --- a/Mathlib/Geometry/RingedSpace/Stalks.lean +++ b/Mathlib/Geometry/RingedSpace/Stalks.lean @@ -57,7 +57,7 @@ section Restrict of `X` at `f x` and the stalk of the restriction of `X` along `f` at t `x`. -/ def restrictStalkIso {U : TopCat} (X : PresheafedSpace.{_, _, v} C) {f : U ⟶ (X : TopCat.{v})} - (h : OpenEmbedding f) (x : U) : (X.restrict h).presheaf.stalk x ≅ X.presheaf.stalk (f x) := + (h : IsOpenEmbedding f) (x : U) : (X.restrict h).presheaf.stalk x ≅ X.presheaf.stalk (f x) := haveI := initial_of_adjunction (h.isOpenMap.adjunctionNhds x) Final.colimitIso (h.isOpenMap.functorNhds x).op ((OpenNhds.inclusion (f x)).op ⋙ X.presheaf) -- As a left adjoint, the functor `h.is_open_map.functor_nhds x` is initial. @@ -67,7 +67,7 @@ def restrictStalkIso {U : TopCat} (X : PresheafedSpace.{_, _, v} C) {f : U ⟶ ( -- Porting note (#11119): removed `simp` attribute, for left hand side is not in simple normal form. @[elementwise, reassoc] theorem restrictStalkIso_hom_eq_germ {U : TopCat} (X : PresheafedSpace.{_, _, v} C) - {f : U ⟶ (X : TopCat.{v})} (h : OpenEmbedding f) (V : Opens U) (x : U) (hx : x ∈ V) : + {f : U ⟶ (X : TopCat.{v})} (h : IsOpenEmbedding f) (V : Opens U) (x : U) (hx : x ∈ V) : (X.restrict h).presheaf.germ _ x hx ≫ (restrictStalkIso X h x).hom = X.presheaf.germ (h.isOpenMap.functor.obj V) (f x) ⟨x, hx, rfl⟩ := colimit.ι_pre ((OpenNhds.inclusion (f x)).op ⋙ X.presheaf) (h.isOpenMap.functorNhds x).op @@ -77,14 +77,14 @@ theorem restrictStalkIso_hom_eq_germ {U : TopCat} (X : PresheafedSpace.{_, _, v} -- as the simpNF linter claims they never apply. @[simp, elementwise, reassoc] theorem restrictStalkIso_inv_eq_germ {U : TopCat} (X : PresheafedSpace.{_, _, v} C) - {f : U ⟶ (X : TopCat.{v})} (h : OpenEmbedding f) (V : Opens U) (x : U) (hx : x ∈ V) : + {f : U ⟶ (X : TopCat.{v})} (h : IsOpenEmbedding f) (V : Opens U) (x : U) (hx : x ∈ V) : X.presheaf.germ (h.isOpenMap.functor.obj V) (f x) ⟨x, hx, rfl⟩ ≫ (restrictStalkIso X h x).inv = (X.restrict h).presheaf.germ _ x hx := by rw [← restrictStalkIso_hom_eq_germ, Category.assoc, Iso.hom_inv_id, Category.comp_id] theorem restrictStalkIso_inv_eq_ofRestrict {U : TopCat} (X : PresheafedSpace.{_, _, v} C) - {f : U ⟶ (X : TopCat.{v})} (h : OpenEmbedding f) (x : U) : + {f : U ⟶ (X : TopCat.{v})} (h : IsOpenEmbedding f) (x : U) : (X.restrictStalkIso h x).inv = (X.ofRestrict h).stalkMap x := by -- We can't use `ext` here due to https://github.com/leanprover/std4/pull/159 refine colimit.hom_ext fun V => ?_ @@ -99,7 +99,7 @@ theorem restrictStalkIso_inv_eq_ofRestrict {U : TopCat} (X : PresheafedSpace.{_, exact (colimit.w ((OpenNhds.inclusion (f x)).op ⋙ X.presheaf) i.op).symm instance ofRestrict_stalkMap_isIso {U : TopCat} (X : PresheafedSpace.{_, _, v} C) - {f : U ⟶ (X : TopCat.{v})} (h : OpenEmbedding f) (x : U) : + {f : U ⟶ (X : TopCat.{v})} (h : IsOpenEmbedding f) (x : U) : IsIso ((X.ofRestrict h).stalkMap x) := by rw [← restrictStalkIso_inv_eq_ofRestrict]; infer_instance diff --git a/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean b/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean index a09f3bc5eb9bd..601ec698d26dd 100644 --- a/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean +++ b/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean @@ -636,10 +636,13 @@ protected theorem ClosedEmbedding.measurableEmbedding {f : α → β} (h : Close MeasurableEmbedding f := h.toEmbedding.measurableEmbedding h.isClosed_range.measurableSet -protected theorem OpenEmbedding.measurableEmbedding {f : α → β} (h : OpenEmbedding f) : +protected theorem IsOpenEmbedding.measurableEmbedding {f : α → β} (h : IsOpenEmbedding f) : MeasurableEmbedding f := h.toEmbedding.measurableEmbedding h.isOpen_range.measurableSet +@[deprecated (since := "2024-10-18")] +alias OpenEmbedding.measurableEmbedding := IsOpenEmbedding.measurableEmbedding + instance Empty.borelSpace : BorelSpace Empty := ⟨borel_eq_top_of_discrete.symm⟩ diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 7268d68dd4275..408d7863a320a 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -190,7 +190,7 @@ lemma eisensteinSeries_tendstoLocallyUniformlyOn {k : ℤ} {N : ℕ} (hk : 3 ≤ apply TendstoLocallyUniformlyOn.comp (s := ⊤) _ _ _ (PartialHomeomorph.continuousOn_symm _) · simp only [SlashInvariantForm.toFun_eq_coe, Set.top_eq_univ, tendstoLocallyUniformlyOn_univ] apply eisensteinSeries_tendstoLocallyUniformly hk - · simp only [OpenEmbedding.toPartialHomeomorph_target, Set.top_eq_univ, mapsTo_range_iff, + · simp only [IsOpenEmbedding.toPartialHomeomorph_target, Set.top_eq_univ, mapsTo_range_iff, Set.mem_univ, forall_const] end summability diff --git a/Mathlib/Topology/Algebra/ConstMulAction.lean b/Mathlib/Topology/Algebra/ConstMulAction.lean index 7f721b2b8dbdd..e8e5e5ff0b50e 100644 --- a/Mathlib/Topology/Algebra/ConstMulAction.lean +++ b/Mathlib/Topology/Algebra/ConstMulAction.lean @@ -258,7 +258,7 @@ theorem subset_interior_smul_right {s : Set G} {t : Set α} : s • interior t @[to_additive (attr := simp)] theorem smul_mem_nhds_smul_iff {t : Set α} (g : G) {a : α} : g • t ∈ 𝓝 (g • a) ↔ t ∈ 𝓝 a := - (Homeomorph.smul g).openEmbedding.image_mem_nhds + (Homeomorph.smul g).isOpenEmbedding.image_mem_nhds @[to_additive] alias ⟨_, smul_mem_nhds_smul⟩ := smul_mem_nhds_smul_iff diff --git a/Mathlib/Topology/Algebra/Constructions/DomMulAct.lean b/Mathlib/Topology/Algebra/Constructions/DomMulAct.lean index 9c399af02e789..a907bb2550a2a 100644 --- a/Mathlib/Topology/Algebra/Constructions/DomMulAct.lean +++ b/Mathlib/Topology/Algebra/Constructions/DomMulAct.lean @@ -49,15 +49,21 @@ theorem coe_mkHomeomorph_symm : ⇑(mkHomeomorph : M ≃ₜ Mᵈᵐᵃ).symm = m @[to_additive] theorem inducing_mk : Inducing (@mk M) := mkHomeomorph.inducing @[to_additive] theorem embedding_mk : Embedding (@mk M) := mkHomeomorph.embedding -@[to_additive] theorem openEmbedding_mk : OpenEmbedding (@mk M) := mkHomeomorph.openEmbedding +@[to_additive] theorem isOpenEmbedding_mk : IsOpenEmbedding (@mk M) := mkHomeomorph.isOpenEmbedding @[to_additive] theorem closedEmbedding_mk : ClosedEmbedding (@mk M) := mkHomeomorph.closedEmbedding @[to_additive] theorem quotientMap_mk : QuotientMap (@mk M) := mkHomeomorph.quotientMap +@[deprecated (since := "2024-10-18")] +alias openEmbedding_mk := isOpenEmbedding_mk + @[to_additive] theorem inducing_mk_symm : Inducing (@mk M).symm := mkHomeomorph.symm.inducing @[to_additive] theorem embedding_mk_symm : Embedding (@mk M).symm := mkHomeomorph.symm.embedding @[to_additive] -theorem openEmbedding_mk_symm : OpenEmbedding (@mk M).symm := mkHomeomorph.symm.openEmbedding +theorem isOpenEmbedding_mk_symm : IsOpenEmbedding (@mk M).symm := mkHomeomorph.symm.isOpenEmbedding + +@[deprecated (since := "2024-10-18")] +alias openEmbedding_mk_symm := isOpenEmbedding_mk_symm @[to_additive] theorem closedEmbedding_mk_symm : ClosedEmbedding (@mk M).symm := mkHomeomorph.symm.closedEmbedding @@ -108,7 +114,7 @@ instance instCompactSpace [CompactSpace M] : CompactSpace Mᵈᵐᵃ := @[to_additive] instance instLocallyCompactSpace [LocallyCompactSpace M] : LocallyCompactSpace Mᵈᵐᵃ := - openEmbedding_mk_symm.locallyCompactSpace + isOpenEmbedding_mk_symm.locallyCompactSpace @[to_additive] instance instWeaklyLocallyCompactSpace [WeaklyLocallyCompactSpace M] : diff --git a/Mathlib/Topology/Algebra/Nonarchimedean/Basic.lean b/Mathlib/Topology/Algebra/Nonarchimedean/Basic.lean index 56b953c25c271..06b301dde6cb3 100644 --- a/Mathlib/Topology/Algebra/Nonarchimedean/Basic.lean +++ b/Mathlib/Topology/Algebra/Nonarchimedean/Basic.lean @@ -60,7 +60,7 @@ variable {K : Type*} [Group K] [TopologicalSpace K] [NonarchimedeanGroup K] /-- If a topological group embeds into a nonarchimedean group, then it is nonarchimedean. -/ @[to_additive] -theorem nonarchimedean_of_emb (f : G →* H) (emb : OpenEmbedding f) : NonarchimedeanGroup H := +theorem nonarchimedean_of_emb (f : G →* H) (emb : IsOpenEmbedding f) : NonarchimedeanGroup H := { is_nonarchimedean := fun U hU => have h₁ : f ⁻¹' U ∈ 𝓝 (1 : G) := by apply emb.continuous.tendsto diff --git a/Mathlib/Topology/Bases.lean b/Mathlib/Topology/Bases.lean index 46d7a2dd5d1df..e953974c7c1dd 100644 --- a/Mathlib/Topology/Bases.lean +++ b/Mathlib/Topology/Bases.lean @@ -887,8 +887,8 @@ theorem IsTopologicalBasis.sum {s : Set (Set α)} (hs : IsTopologicalBasis s) {t IsTopologicalBasis ((fun u => Sum.inl '' u) '' s ∪ (fun u => Sum.inr '' u) '' t) := by apply isTopologicalBasis_of_isOpen_of_nhds · rintro u (⟨w, hw, rfl⟩ | ⟨w, hw, rfl⟩) - · exact openEmbedding_inl.isOpenMap w (hs.isOpen hw) - · exact openEmbedding_inr.isOpenMap w (ht.isOpen hw) + · exact isOpenEmbedding_inl.isOpenMap w (hs.isOpen hw) + · exact isOpenEmbedding_inr.isOpenMap w (ht.isOpen hw) · rintro (x | x) u hxu u_open · obtain ⟨v, vs, xv, vu⟩ : ∃ v ∈ s, x ∈ v ∧ v ⊆ Sum.inl ⁻¹' u := hs.exists_subset_of_mem_open hxu (isOpen_sum_iff.1 u_open).1 diff --git a/Mathlib/Topology/Category/CompHausLike/Limits.lean b/Mathlib/Topology/Category/CompHausLike/Limits.lean index 845abc60f56e7..3e5a965f2dce5 100644 --- a/Mathlib/Topology/Category/CompHausLike/Limits.lean +++ b/Mathlib/Topology/Category/CompHausLike/Limits.lean @@ -150,20 +150,26 @@ variable {P : TopCat.{u} → Prop} [HasExplicitFiniteCoproducts.{0} P] example : HasFiniteCoproducts (CompHausLike.{u} P) := inferInstance /-- The inclusion maps into the explicit finite coproduct are open embeddings. -/ -lemma finiteCoproduct.openEmbedding_ι (a : α) : - OpenEmbedding (finiteCoproduct.ι X a) := - openEmbedding_sigmaMk (σ := fun a ↦ (X a)) +lemma finiteCoproduct.isOpenEmbedding_ι (a : α) : + IsOpenEmbedding (finiteCoproduct.ι X a) := + isOpenEmbedding_sigmaMk (σ := fun a ↦ (X a)) + +@[deprecated (since := "2024-10-18")] +alias finiteCoproduct.openEmbedding_ι := finiteCoproduct.isOpenEmbedding_ι /-- The inclusion maps into the abstract finite coproduct are open embeddings. -/ -lemma Sigma.openEmbedding_ι (a : α) : - OpenEmbedding (Sigma.ι X a) := by - refine OpenEmbedding.of_comp _ (homeoOfIso ((colimit.isColimit _).coconePointUniqueUpToIso - (finiteCoproduct.isColimit X))).openEmbedding ?_ - convert finiteCoproduct.openEmbedding_ι X a +lemma Sigma.isOpenEmbedding_ι (a : α) : + IsOpenEmbedding (Sigma.ι X a) := by + refine IsOpenEmbedding.of_comp _ (homeoOfIso ((colimit.isColimit _).coconePointUniqueUpToIso + (finiteCoproduct.isColimit X))).isOpenEmbedding ?_ + convert finiteCoproduct.isOpenEmbedding_ι X a ext x change (Sigma.ι X a ≫ _) x = _ simp +@[deprecated (since := "2024-10-18")] +alias Sigma.openEmbedding_ι := Sigma.isOpenEmbedding_ι + /-- The functor to `TopCat` preserves finite coproducts if they exist. -/ instance (P) [HasExplicitFiniteCoproducts.{0} P] : PreservesFiniteCoproducts (compHausLikeToTop P) := by @@ -332,12 +338,12 @@ instance [HasExplicitPullbacksOfInclusions P] : HasPullbacksOfInclusions (CompHa theorem hasPullbacksOfInclusions (hP' : ∀ ⦃X Y B : CompHausLike.{u} P⦄ (f : X ⟶ B) (g : Y ⟶ B) - (_ : OpenEmbedding f), HasExplicitPullback f g) : + (_ : IsOpenEmbedding f), HasExplicitPullback f g) : HasExplicitPullbacksOfInclusions P := { hasProp := by intro _ _ _ f apply hP' - exact Sigma.openEmbedding_ι _ _ } + exact Sigma.isOpenEmbedding_ι _ _ } /-- The functor to `TopCat` preserves pullbacks of inclusions if they exist. -/ noncomputable instance [HasExplicitPullbacksOfInclusions P] : @@ -350,7 +356,7 @@ instance [HasExplicitPullbacksOfInclusions P] : FinitaryExtensive (CompHausLike finitaryExtensive_of_preserves_and_reflects (compHausLikeToTop P) theorem finitaryExtensive (hP' : ∀ ⦃X Y B : CompHausLike.{u} P⦄ (f : X ⟶ B) (g : Y ⟶ B) - (_ : OpenEmbedding f), HasExplicitPullback f g) : + (_ : IsOpenEmbedding f), HasExplicitPullback f g) : FinitaryExtensive (CompHausLike P) := have := hasPullbacksOfInclusions hP' finitaryExtensive_of_preserves_and_reflects (compHausLikeToTop P) diff --git a/Mathlib/Topology/Category/Stonean/Limits.lean b/Mathlib/Topology/Category/Stonean/Limits.lean index 9a344de319c22..ed7151753cb73 100644 --- a/Mathlib/Topology/Category/Stonean/Limits.lean +++ b/Mathlib/Topology/Category/Stonean/Limits.lean @@ -24,7 +24,7 @@ namespace Stonean instance : HasExplicitFiniteCoproducts.{w, u} (fun Y ↦ ExtremallyDisconnected Y) where hasProp _ := { hasProp := show ExtremallyDisconnected (Σ (_a : _), _) from inferInstance} -variable {X Y Z : Stonean} {f : X ⟶ Z} (i : Y ⟶ Z) (hi : OpenEmbedding f) +variable {X Y Z : Stonean} {f : X ⟶ Z} (i : Y ⟶ Z) (hi : IsOpenEmbedding f) include hi lemma extremallyDisconnected_preimage : ExtremallyDisconnected (i ⁻¹' (Set.range f)) where @@ -35,7 +35,7 @@ lemma extremallyDisconnected_preimage : ExtremallyDisconnected (i ⁻¹' (Set.ra rw [← (closure U).preimage_image_eq Subtype.coe_injective, ← h.1.closedEmbedding_subtype_val.closure_image_eq U] exact isOpen_induced (ExtremallyDisconnected.open_closure _ - (h.2.openEmbedding_subtype_val.isOpenMap U hU)) + (h.2.isOpenEmbedding_subtypeVal.isOpenMap U hU)) lemma extremallyDisconnected_pullback : ExtremallyDisconnected {xy : X × Y | f xy.1 = i xy.2} := have := extremallyDisconnected_preimage i hi diff --git a/Mathlib/Topology/Category/TopCat/Basic.lean b/Mathlib/Topology/Category/TopCat/Basic.lean index f1a2371e9013c..57a2905802c66 100644 --- a/Mathlib/Topology/Category/TopCat/Basic.lean +++ b/Mathlib/Topology/Category/TopCat/Basic.lean @@ -168,29 +168,41 @@ lemma isIso_of_bijective_of_isClosedMap {X Y : TopCat.{u}} (f : X ⟶ Y) inferInstanceAs <| IsIso (TopCat.isoOfHomeo e).hom -- Porting note: simpNF requested partially simped version below -theorem openEmbedding_iff_comp_isIso {X Y Z : TopCat} (f : X ⟶ Y) (g : Y ⟶ Z) [IsIso g] : - OpenEmbedding (f ≫ g) ↔ OpenEmbedding f := - (TopCat.homeoOfIso (asIso g)).openEmbedding.of_comp_iff f +theorem isOpenEmbedding_iff_comp_isIso {X Y Z : TopCat} (f : X ⟶ Y) (g : Y ⟶ Z) [IsIso g] : + IsOpenEmbedding (f ≫ g) ↔ IsOpenEmbedding f := + (TopCat.homeoOfIso (asIso g)).isOpenEmbedding.of_comp_iff f + +@[deprecated (since := "2024-10-18")] +alias openEmbedding_iff_comp_isIso := isOpenEmbedding_iff_comp_isIso @[simp] -theorem openEmbedding_iff_comp_isIso' {X Y Z : TopCat} (f : X ⟶ Y) (g : Y ⟶ Z) [IsIso g] : - OpenEmbedding ((forget TopCat).map f ≫ (forget TopCat).map g) ↔ OpenEmbedding f := by +theorem isOpenEmbedding_iff_comp_isIso' {X Y Z : TopCat} (f : X ⟶ Y) (g : Y ⟶ Z) [IsIso g] : + IsOpenEmbedding ((forget TopCat).map f ≫ (forget TopCat).map g) ↔ IsOpenEmbedding f := by simp only [← Functor.map_comp] - exact openEmbedding_iff_comp_isIso f g + exact isOpenEmbedding_iff_comp_isIso f g + +@[deprecated (since := "2024-10-18")] +alias openEmbedding_iff_comp_isIso' := isOpenEmbedding_iff_comp_isIso' -- Porting note: simpNF requested partially simped version below -theorem openEmbedding_iff_isIso_comp {X Y Z : TopCat} (f : X ⟶ Y) (g : Y ⟶ Z) [IsIso f] : - OpenEmbedding (f ≫ g) ↔ OpenEmbedding g := by +theorem isOpenEmbedding_iff_isIso_comp {X Y Z : TopCat} (f : X ⟶ Y) (g : Y ⟶ Z) [IsIso f] : + IsOpenEmbedding (f ≫ g) ↔ IsOpenEmbedding g := by constructor · intro h - convert h.comp (TopCat.homeoOfIso (asIso f).symm).openEmbedding + convert h.comp (TopCat.homeoOfIso (asIso f).symm).isOpenEmbedding exact congrArg _ (IsIso.inv_hom_id_assoc f g).symm - · exact fun h => h.comp (TopCat.homeoOfIso (asIso f)).openEmbedding + · exact fun h => h.comp (TopCat.homeoOfIso (asIso f)).isOpenEmbedding + +@[deprecated (since := "2024-10-18")] +alias openEmbedding_iff_isIso_comp := isOpenEmbedding_iff_isIso_comp @[simp] -theorem openEmbedding_iff_isIso_comp' {X Y Z : TopCat} (f : X ⟶ Y) (g : Y ⟶ Z) [IsIso f] : - OpenEmbedding ((forget TopCat).map f ≫ (forget TopCat).map g) ↔ OpenEmbedding g := by +theorem isOpenEmbedding_iff_isIso_comp' {X Y Z : TopCat} (f : X ⟶ Y) (g : Y ⟶ Z) [IsIso f] : + IsOpenEmbedding ((forget TopCat).map f ≫ (forget TopCat).map g) ↔ IsOpenEmbedding g := by simp only [← Functor.map_comp] - exact openEmbedding_iff_isIso_comp f g + exact isOpenEmbedding_iff_isIso_comp f g + +@[deprecated (since := "2024-10-18")] +alias openEmbedding_iff_isIso_comp' := isOpenEmbedding_iff_isIso_comp' end TopCat diff --git a/Mathlib/Topology/Category/TopCat/Limits/Products.lean b/Mathlib/Topology/Category/TopCat/Limits/Products.lean index cd45197d42a40..132e2144e2090 100644 --- a/Mathlib/Topology/Category/TopCat/Limits/Products.lean +++ b/Mathlib/Topology/Category/TopCat/Limits/Products.lean @@ -283,7 +283,8 @@ def binaryCofanIsColimit (X Y : TopCat.{u}) : IsColimit (TopCat.binaryCofan X Y) theorem binaryCofan_isColimit_iff {X Y : TopCat} (c : BinaryCofan X Y) : Nonempty (IsColimit c) ↔ - OpenEmbedding c.inl ∧ OpenEmbedding c.inr ∧ IsCompl (Set.range c.inl) (Set.range c.inr) := by + IsOpenEmbedding c.inl ∧ IsOpenEmbedding c.inr ∧ + IsCompl (Set.range c.inl) (Set.range c.inr) := by classical constructor · rintro ⟨h⟩ @@ -293,9 +294,9 @@ theorem binaryCofan_isColimit_iff {X Y : TopCat} (c : BinaryCofan X Y) : h.comp_coconePointUniqueUpToIso_inv (binaryCofanIsColimit X Y) ⟨WalkingPair.right⟩] dsimp refine ⟨(homeoOfIso <| h.coconePointUniqueUpToIso - (binaryCofanIsColimit X Y)).symm.openEmbedding.comp openEmbedding_inl, + (binaryCofanIsColimit X Y)).symm.isOpenEmbedding.comp isOpenEmbedding_inl, (homeoOfIso <| h.coconePointUniqueUpToIso - (binaryCofanIsColimit X Y)).symm.openEmbedding.comp openEmbedding_inr, ?_⟩ + (binaryCofanIsColimit X Y)).symm.isOpenEmbedding.comp isOpenEmbedding_inr, ?_⟩ erw [Set.range_comp, ← eq_compl_iff_isCompl, Set.range_comp _ Sum.inr, ← Set.image_compl_eq (homeoOfIso <| h.coconePointUniqueUpToIso (binaryCofanIsColimit X Y)).symm.bijective, Set.compl_range_inr, Set.image_comp] diff --git a/Mathlib/Topology/Category/TopCat/Limits/Pullbacks.lean b/Mathlib/Topology/Category/TopCat/Limits/Pullbacks.lean index 7427588e292b7..e51d68906803f 100644 --- a/Mathlib/Topology/Category/TopCat/Limits/Pullbacks.lean +++ b/Mathlib/Topology/Category/TopCat/Limits/Pullbacks.lean @@ -284,10 +284,10 @@ W ⟶ Y X ⟶ Z ``` -/ -theorem pullback_map_openEmbedding_of_open_embeddings {W X Y Z S T : TopCat.{u}} (f₁ : W ⟶ S) - (f₂ : X ⟶ S) (g₁ : Y ⟶ T) (g₂ : Z ⟶ T) {i₁ : W ⟶ Y} {i₂ : X ⟶ Z} (H₁ : OpenEmbedding i₁) - (H₂ : OpenEmbedding i₂) (i₃ : S ⟶ T) [H₃ : Mono i₃] (eq₁ : f₁ ≫ i₃ = i₁ ≫ g₁) - (eq₂ : f₂ ≫ i₃ = i₂ ≫ g₂) : OpenEmbedding (pullback.map f₁ f₂ g₁ g₂ i₁ i₂ i₃ eq₁ eq₂) := by +theorem pullback_map_isOpenEmbedding {W X Y Z S T : TopCat.{u}} (f₁ : W ⟶ S) + (f₂ : X ⟶ S) (g₁ : Y ⟶ T) (g₂ : Z ⟶ T) {i₁ : W ⟶ Y} {i₂ : X ⟶ Z} (H₁ : IsOpenEmbedding i₁) + (H₂ : IsOpenEmbedding i₂) (i₃ : S ⟶ T) [H₃ : Mono i₃] (eq₁ : f₁ ≫ i₃ = i₁ ≫ g₁) + (eq₂ : f₂ ≫ i₃ = i₂ ≫ g₂) : IsOpenEmbedding (pullback.map f₁ f₂ g₁ g₂ i₁ i₂ i₃ eq₁ eq₂) := by constructor · apply pullback_map_embedding_of_embeddings f₁ f₂ g₁ g₂ H₁.toEmbedding H₂.toEmbedding i₃ eq₁ eq₂ @@ -298,6 +298,9 @@ theorem pullback_map_openEmbedding_of_open_embeddings {W X Y Z S T : TopCat.{u}} · apply ContinuousMap.continuous_toFun · exact H₂.isOpen_range +@[deprecated (since := "2024-10-18")] +alias pullback_map_openEmbedding_of_open_embeddings := pullback_map_isOpenEmbedding + theorem snd_embedding_of_left_embedding {X Y S : TopCat} {f : X ⟶ S} (H : Embedding f) (g : Y ⟶ S) : Embedding <| ⇑(pullback.snd f g) := by convert (homeoOfIso (asIso (pullback.snd (𝟙 S) g))).embedding.comp @@ -320,30 +323,39 @@ theorem embedding_of_pullback_embeddings {X Y S : TopCat} {f : X ⟶ S} {g : Y rw [← coe_comp, ← limit.w _ WalkingCospan.Hom.inr] rfl -theorem snd_openEmbedding_of_left_openEmbedding {X Y S : TopCat} {f : X ⟶ S} (H : OpenEmbedding f) - (g : Y ⟶ S) : OpenEmbedding <| ⇑(pullback.snd f g) := by - convert (homeoOfIso (asIso (pullback.snd (𝟙 S) g))).openEmbedding.comp - (pullback_map_openEmbedding_of_open_embeddings (i₂ := 𝟙 Y) f g (𝟙 _) g H - (homeoOfIso (Iso.refl _)).openEmbedding (𝟙 _) rfl (by simp)) +theorem snd_isOpenEmbedding_of_left {X Y S : TopCat} {f : X ⟶ S} (H : IsOpenEmbedding f) + (g : Y ⟶ S) : IsOpenEmbedding <| ⇑(pullback.snd f g) := by + convert (homeoOfIso (asIso (pullback.snd (𝟙 S) g))).isOpenEmbedding.comp + (pullback_map_isOpenEmbedding (i₂ := 𝟙 Y) f g (𝟙 _) g H + (homeoOfIso (Iso.refl _)).isOpenEmbedding (𝟙 _) rfl (by simp)) erw [← coe_comp] simp -theorem fst_openEmbedding_of_right_openEmbedding {X Y S : TopCat} (f : X ⟶ S) {g : Y ⟶ S} - (H : OpenEmbedding g) : OpenEmbedding <| ⇑(pullback.fst f g) := by - convert (homeoOfIso (asIso (pullback.fst f (𝟙 S)))).openEmbedding.comp - (pullback_map_openEmbedding_of_open_embeddings (i₁ := 𝟙 X) f g f (𝟙 _) - (homeoOfIso (Iso.refl _)).openEmbedding H (𝟙 _) rfl (by simp)) +@[deprecated (since := "2024-10-18")] +alias snd_openEmbedding_of_left_openEmbedding := snd_isOpenEmbedding_of_left + +theorem fst_isOpenEmbedding_of_right {X Y S : TopCat} (f : X ⟶ S) {g : Y ⟶ S} + (H : IsOpenEmbedding g) : IsOpenEmbedding <| ⇑(pullback.fst f g) := by + convert (homeoOfIso (asIso (pullback.fst f (𝟙 S)))).isOpenEmbedding.comp + (pullback_map_isOpenEmbedding (i₁ := 𝟙 X) f g f (𝟙 _) + (homeoOfIso (Iso.refl _)).isOpenEmbedding H (𝟙 _) rfl (by simp)) erw [← coe_comp] simp +@[deprecated (since := "2024-10-18")] +alias fst_openEmbedding_of_right_openEmbedding := fst_isOpenEmbedding_of_right + /-- If `X ⟶ S`, `Y ⟶ S` are open embeddings, then so is `X ×ₛ Y ⟶ S`. -/ -theorem openEmbedding_of_pullback_open_embeddings {X Y S : TopCat} {f : X ⟶ S} {g : Y ⟶ S} - (H₁ : OpenEmbedding f) (H₂ : OpenEmbedding g) : - OpenEmbedding (limit.π (cospan f g) WalkingCospan.one) := by - convert H₂.comp (snd_openEmbedding_of_left_openEmbedding H₁ g) +theorem isOpenEmbedding_of_pullback_open_embeddings {X Y S : TopCat} {f : X ⟶ S} {g : Y ⟶ S} + (H₁ : IsOpenEmbedding f) (H₂ : IsOpenEmbedding g) : + IsOpenEmbedding (limit.π (cospan f g) WalkingCospan.one) := by + convert H₂.comp (snd_isOpenEmbedding_of_left H₁ g) rw [← coe_comp, ← limit.w _ WalkingCospan.Hom.inr] rfl +@[deprecated (since := "2024-10-18")] +alias openEmbedding_of_pullback_open_embeddings := isOpenEmbedding_of_pullback_open_embeddings + theorem fst_iso_of_right_embedding_range_subset {X Y S : TopCat} (f : X ⟶ S) {g : Y ⟶ S} (hg : Embedding g) (H : Set.range f ⊆ Set.range g) : IsIso (pullback.fst f g) := by diff --git a/Mathlib/Topology/Category/TopCat/OpenNhds.lean b/Mathlib/Topology/Category/TopCat/OpenNhds.lean index 863d3a8c4d612..fdfe88f737c64 100644 --- a/Mathlib/Topology/Category/TopCat/OpenNhds.lean +++ b/Mathlib/Topology/Category/TopCat/OpenNhds.lean @@ -86,8 +86,11 @@ def inclusion (x : X) : OpenNhds x ⥤ Opens X := theorem inclusion_obj (x : X) (U) (p) : (inclusion x).obj ⟨U, p⟩ = U := rfl -theorem openEmbedding {x : X} (U : OpenNhds x) : OpenEmbedding U.1.inclusion' := - U.1.openEmbedding +theorem isOpenEmbedding {x : X} (U : OpenNhds x) : IsOpenEmbedding U.1.inclusion' := + U.1.isOpenEmbedding + +@[deprecated (since := "2024-10-18")] +alias openEmbedding := isOpenEmbedding /-- The preimage functor from neighborhoods of `f x` to neighborhoods of `x`. -/ def map (x : X) : OpenNhds (f x) ⥤ OpenNhds x where diff --git a/Mathlib/Topology/Category/TopCat/Opens.lean b/Mathlib/Topology/Category/TopCat/Opens.lean index eb8eb48fb88be..de97a8d22ea44 100644 --- a/Mathlib/Topology/Category/TopCat/Opens.lean +++ b/Mathlib/Topology/Category/TopCat/Opens.lean @@ -119,8 +119,11 @@ def inclusion' {X : TopCat.{u}} (U : Opens X) : (toTopCat X).obj U ⟶ X where theorem coe_inclusion' {X : TopCat} {U : Opens X} : (inclusion' U : U → X) = Subtype.val := rfl -theorem openEmbedding {X : TopCat.{u}} (U : Opens X) : OpenEmbedding (inclusion' U) := - IsOpen.openEmbedding_subtype_val U.2 +theorem isOpenEmbedding {X : TopCat.{u}} (U : Opens X) : IsOpenEmbedding (inclusion' U) := + IsOpen.isOpenEmbedding_subtypeVal U.2 + +@[deprecated (since := "2024-10-18")] +alias openEmbedding := isOpenEmbedding /-- The inclusion of the top open subset (i.e. the whole space) is an isomorphism. -/ @@ -294,20 +297,26 @@ instance IsOpenMap.functorFullOfMono {X Y : TopCat} {f : X ⟶ Y} (hf : IsOpenMa instance IsOpenMap.functor_faithful {X Y : TopCat} {f : X ⟶ Y} (hf : IsOpenMap f) : hf.functor.Faithful where -lemma OpenEmbedding.functor_obj_injective {X Y : TopCat} {f : X ⟶ Y} (hf : OpenEmbedding f) : +lemma IsOpenEmbedding.functor_obj_injective {X Y : TopCat} {f : X ⟶ Y} (hf : IsOpenEmbedding f) : Function.Injective hf.isOpenMap.functor.obj := fun _ _ e ↦ Opens.ext (Set.image_injective.mpr hf.inj (congr_arg (↑· : Opens Y → Set Y) e)) +@[deprecated (since := "2024-10-18")] +alias OpenEmbedding.functor_obj_injective := IsOpenEmbedding.functor_obj_injective + namespace TopologicalSpace.Opens open TopologicalSpace @[simp] -theorem openEmbedding_obj_top {X : TopCat} (U : Opens X) : - U.openEmbedding.isOpenMap.functor.obj ⊤ = U := by +theorem isOpenEmbedding_obj_top {X : TopCat} (U : Opens X) : + U.isOpenEmbedding.isOpenMap.functor.obj ⊤ = U := by ext1 exact Set.image_univ.trans Subtype.range_coe +@[deprecated (since := "2024-10-18")] +alias openEmbedding_obj_top := isOpenEmbedding_obj_top + @[simp] theorem inclusion'_map_eq_top {X : TopCat} (U : Opens X) : (Opens.map U.inclusion').obj U = ⊤ := by ext1 @@ -315,10 +324,10 @@ theorem inclusion'_map_eq_top {X : TopCat} (U : Opens X) : (Opens.map U.inclusio @[simp] theorem adjunction_counit_app_self {X : TopCat} (U : Opens X) : - U.openEmbedding.isOpenMap.adjunction.counit.app U = eqToHom (by simp) := Subsingleton.elim _ _ + U.isOpenEmbedding.isOpenMap.adjunction.counit.app U = eqToHom (by simp) := Subsingleton.elim _ _ theorem inclusion'_top_functor (X : TopCat) : - (@Opens.openEmbedding X ⊤).isOpenMap.functor = map (inclusionTopIso X).inv := by + (@Opens.isOpenEmbedding X ⊤).isOpenMap.functor = map (inclusionTopIso X).inv := by refine CategoryTheory.Functor.ext ?_ ?_ · intro U ext x @@ -347,23 +356,23 @@ lemma set_range_inclusion' {X : TopCat} (U : Opens X) : @[simp] theorem functor_map_eq_inf {X : TopCat} (U V : Opens X) : - U.openEmbedding.isOpenMap.functor.obj ((Opens.map U.inclusion').obj V) = V ⊓ U := by + U.isOpenEmbedding.isOpenMap.functor.obj ((Opens.map U.inclusion').obj V) = V ⊓ U := by ext1 simp only [IsOpenMap.functor_obj_coe, map_coe, coe_inf, Set.image_preimage_eq_inter_range, set_range_inclusion' U] -theorem map_functor_eq' {X U : TopCat} (f : U ⟶ X) (hf : OpenEmbedding f) (V) : +theorem map_functor_eq' {X U : TopCat} (f : U ⟶ X) (hf : IsOpenEmbedding f) (V) : ((Opens.map f).obj <| hf.isOpenMap.functor.obj V) = V := Opens.ext <| Set.preimage_image_eq _ hf.inj @[simp] theorem map_functor_eq {X : TopCat} {U : Opens X} (V : Opens U) : - ((Opens.map U.inclusion').obj <| U.openEmbedding.isOpenMap.functor.obj V) = V := - TopologicalSpace.Opens.map_functor_eq' _ U.openEmbedding V + ((Opens.map U.inclusion').obj <| U.isOpenEmbedding.isOpenMap.functor.obj V) = V := + TopologicalSpace.Opens.map_functor_eq' _ U.isOpenEmbedding V @[simp] theorem adjunction_counit_map_functor {X : TopCat} {U : Opens X} (V : Opens U) : - U.openEmbedding.isOpenMap.adjunction.counit.app (U.openEmbedding.isOpenMap.functor.obj V) = + U.isOpenEmbedding.isOpenMap.adjunction.counit.app (U.isOpenEmbedding.isOpenMap.functor.obj V) = eqToHom (by dsimp; rw [map_functor_eq V]) := by subsingleton diff --git a/Mathlib/Topology/Clopen.lean b/Mathlib/Topology/Clopen.lean index ff10a04698b33..511935a8e3b98 100644 --- a/Mathlib/Topology/Clopen.lean +++ b/Mathlib/Topology/Clopen.lean @@ -118,7 +118,7 @@ theorem isClopen_range_inr : IsClopen (range (Sum.inr : Y → X ⊕ Y)) := theorem isClopen_range_sigmaMk {X : ι → Type*} [∀ i, TopologicalSpace (X i)] {i : ι} : IsClopen (Set.range (@Sigma.mk ι X i)) := - ⟨closedEmbedding_sigmaMk.isClosed_range, openEmbedding_sigmaMk.isOpen_range⟩ + ⟨closedEmbedding_sigmaMk.isClosed_range, isOpenEmbedding_sigmaMk.isOpen_range⟩ protected theorem QuotientMap.isClopen_preimage {f : X → Y} (hf : QuotientMap f) {s : Set Y} : IsClopen (f ⁻¹' s) ↔ IsClopen s := diff --git a/Mathlib/Topology/Compactification/OnePoint.lean b/Mathlib/Topology/Compactification/OnePoint.lean index 622f86227bd24..e870661f60d4f 100644 --- a/Mathlib/Topology/Compactification/OnePoint.lean +++ b/Mathlib/Topology/Compactification/OnePoint.lean @@ -241,28 +241,31 @@ theorem continuous_coe : Continuous ((↑) : X → OnePoint X) := theorem isOpenMap_coe : IsOpenMap ((↑) : X → OnePoint X) := fun _ => isOpen_image_coe.2 -theorem openEmbedding_coe : OpenEmbedding ((↑) : X → OnePoint X) := - openEmbedding_of_continuous_injective_open continuous_coe coe_injective isOpenMap_coe +theorem isOpenEmbedding_coe : IsOpenEmbedding ((↑) : X → OnePoint X) := + isOpenEmbedding_of_continuous_injective_open continuous_coe coe_injective isOpenMap_coe + +@[deprecated (since := "2024-10-18")] +alias openEmbedding_coe := isOpenEmbedding_coe theorem isOpen_range_coe : IsOpen (range ((↑) : X → OnePoint X)) := - openEmbedding_coe.isOpen_range + isOpenEmbedding_coe.isOpen_range theorem isClosed_infty : IsClosed ({∞} : Set (OnePoint X)) := by rw [← compl_range_coe, isClosed_compl_iff] exact isOpen_range_coe theorem nhds_coe_eq (x : X) : 𝓝 ↑x = map ((↑) : X → OnePoint X) (𝓝 x) := - (openEmbedding_coe.map_nhds_eq x).symm + (isOpenEmbedding_coe.map_nhds_eq x).symm theorem nhdsWithin_coe_image (s : Set X) (x : X) : 𝓝[(↑) '' s] (x : OnePoint X) = map (↑) (𝓝[s] x) := - (openEmbedding_coe.toEmbedding.map_nhdsWithin_eq _ _).symm + (isOpenEmbedding_coe.toEmbedding.map_nhdsWithin_eq _ _).symm theorem nhdsWithin_coe (s : Set (OnePoint X)) (x : X) : 𝓝[s] ↑x = map (↑) (𝓝[(↑) ⁻¹' s] x) := - (openEmbedding_coe.map_nhdsWithin_preimage_eq _ _).symm + (isOpenEmbedding_coe.map_nhdsWithin_preimage_eq _ _).symm theorem comap_coe_nhds (x : X) : comap ((↑) : X → OnePoint X) (𝓝 x) = 𝓝 x := - (openEmbedding_coe.toInducing.nhds_eq_comap x).symm + (isOpenEmbedding_coe.toInducing.nhds_eq_comap x).symm /-- If `x` is not an isolated point of `X`, then `x : OnePoint X` is not an isolated point of `OnePoint X`. -/ @@ -427,18 +430,18 @@ theorem denseRange_coe [NoncompactSpace X] : DenseRange ((↑) : X → OnePoint exact dense_compl_singleton _ theorem isDenseEmbedding_coe [NoncompactSpace X] : IsDenseEmbedding ((↑) : X → OnePoint X) := - { openEmbedding_coe with dense := denseRange_coe } + { isOpenEmbedding_coe with dense := denseRange_coe } @[deprecated (since := "2024-09-30")] alias denseEmbedding_coe := isDenseEmbedding_coe @[simp, norm_cast] theorem specializes_coe {x y : X} : (x : OnePoint X) ⤳ y ↔ x ⤳ y := - openEmbedding_coe.toInducing.specializes_iff + isOpenEmbedding_coe.toInducing.specializes_iff @[simp, norm_cast] theorem inseparable_coe {x y : X} : Inseparable (x : OnePoint X) y ↔ Inseparable x y := - openEmbedding_coe.toInducing.inseparable_iff + isOpenEmbedding_coe.toInducing.inseparable_iff theorem not_specializes_infty_coe {x : X} : ¬Specializes ∞ (x : OnePoint X) := isClosed_infty.not_specializes rfl (coe_ne_infty x) diff --git a/Mathlib/Topology/Compactness/LocallyCompact.lean b/Mathlib/Topology/Compactness/LocallyCompact.lean index 8d10e1ce3da2e..5bba2a34cb127 100644 --- a/Mathlib/Topology/Compactness/LocallyCompact.lean +++ b/Mathlib/Topology/Compactness/LocallyCompact.lean @@ -199,10 +199,13 @@ protected theorem ClosedEmbedding.locallyCompactSpace [LocallyCompactSpace Y] {f (hf : ClosedEmbedding f) : LocallyCompactSpace X := hf.toInducing.locallyCompactSpace hf.isClosed_range.isLocallyClosed -protected theorem OpenEmbedding.locallyCompactSpace [LocallyCompactSpace Y] {f : X → Y} - (hf : OpenEmbedding f) : LocallyCompactSpace X := +protected theorem IsOpenEmbedding.locallyCompactSpace [LocallyCompactSpace Y] {f : X → Y} + (hf : IsOpenEmbedding f) : LocallyCompactSpace X := hf.toInducing.locallyCompactSpace hf.isOpen_range.isLocallyClosed +@[deprecated (since := "2024-10-18")] +alias OpenEmbedding.locallyCompactSpace := IsOpenEmbedding.locallyCompactSpace + protected theorem IsLocallyClosed.locallyCompactSpace [LocallyCompactSpace X] {s : Set X} (hs : IsLocallyClosed s) : LocallyCompactSpace s := embedding_subtype_val.locallyCompactSpace <| by rwa [Subtype.range_val] diff --git a/Mathlib/Topology/Connected/LocallyConnected.lean b/Mathlib/Topology/Connected/LocallyConnected.lean index 52d6a10719be2..065d40db50471 100644 --- a/Mathlib/Topology/Connected/LocallyConnected.lean +++ b/Mathlib/Topology/Connected/LocallyConnected.lean @@ -120,8 +120,8 @@ theorem locallyConnectedSpace_of_connected_bases {ι : Type*} (b : α → ι → (fun i hi => ⟨b x i, ⟨(hbasis x).mem_of_mem hi, hconnected x i hi⟩, subset_rfl⟩) fun s hs => ⟨(hbasis x).index s hs.1, ⟨(hbasis x).property_index hs.1, (hbasis x).set_index_subset hs.1⟩⟩ -theorem OpenEmbedding.locallyConnectedSpace [LocallyConnectedSpace α] [TopologicalSpace β] - {f : β → α} (h : OpenEmbedding f) : LocallyConnectedSpace β := by +theorem IsOpenEmbedding.locallyConnectedSpace [LocallyConnectedSpace α] [TopologicalSpace β] + {f : β → α} (h : IsOpenEmbedding f) : LocallyConnectedSpace β := by refine locallyConnectedSpace_of_connected_bases (fun _ s ↦ f ⁻¹' s) (fun x s ↦ (IsOpen s ∧ f x ∈ s ∧ IsConnected s) ∧ s ⊆ range f) (fun x ↦ ?_) (fun x s hxs ↦ hxs.1.2.2.isPreconnected.preimage_of_isOpenMap h.inj h.isOpenMap hxs.2) @@ -129,8 +129,11 @@ theorem OpenEmbedding.locallyConnectedSpace [LocallyConnectedSpace α] [Topologi exact LocallyConnectedSpace.open_connected_basis (f x) |>.restrict_subset (h.isOpen_range.mem_nhds <| mem_range_self _) |>.comap _ +@[deprecated (since := "2024-10-18")] +alias OpenEmbedding.locallyConnectedSpace := IsOpenEmbedding.locallyConnectedSpace + theorem IsOpen.locallyConnectedSpace [LocallyConnectedSpace α] {U : Set α} (hU : IsOpen U) : LocallyConnectedSpace U := - hU.openEmbedding_subtype_val.locallyConnectedSpace + hU.isOpenEmbedding_subtypeVal.locallyConnectedSpace end LocallyConnectedSpace diff --git a/Mathlib/Topology/Connected/PathConnected.lean b/Mathlib/Topology/Connected/PathConnected.lean index 93ef11a0f5e30..3a8377f938164 100644 --- a/Mathlib/Topology/Connected/PathConnected.lean +++ b/Mathlib/Topology/Connected/PathConnected.lean @@ -1154,7 +1154,7 @@ theorem pathConnected_subset_basis {U : Set X} (h : IsOpen U) (hx : x ∈ U) : (𝓝 x).HasBasis (fun s : Set X => s ∈ 𝓝 x ∧ IsPathConnected s ∧ s ⊆ U) id := (path_connected_basis x).hasBasis_self_subset (IsOpen.mem_nhds h hx) -theorem OpenEmbedding.locPathConnectedSpace {e : Y → X} (he : OpenEmbedding e) : +theorem IsOpenEmbedding.locPathConnectedSpace {e : Y → X} (he : IsOpenEmbedding e) : LocPathConnectedSpace Y := have (y : Y) : (𝓝 y).HasBasis (fun s ↦ s ∈ 𝓝 (e y) ∧ IsPathConnected s ∧ s ⊆ range e) (e ⁻¹' ·) := @@ -1162,8 +1162,11 @@ theorem OpenEmbedding.locPathConnectedSpace {e : Y → X} (he : OpenEmbedding e) .of_bases this fun x s ⟨_, hs, hse⟩ ↦ by rwa [he.isPathConnected_iff, image_preimage_eq_of_subset hse] +@[deprecated (since := "2024-10-18")] +alias OpenEmbedding.locPathConnectedSpace := IsOpenEmbedding.locPathConnectedSpace + theorem IsOpen.locPathConnectedSpace {U : Set X} (h : IsOpen U) : LocPathConnectedSpace U := - (openEmbedding_subtype_val h).locPathConnectedSpace + (isOpenEmbedding_subtypeVal h).locPathConnectedSpace @[deprecated (since := "2024-10-17")] alias locPathConnected_of_isOpen := IsOpen.locPathConnectedSpace diff --git a/Mathlib/Topology/Constructions.lean b/Mathlib/Topology/Constructions.lean index 82bce23e4c0c9..85297dd644e16 100644 --- a/Mathlib/Topology/Constructions.lean +++ b/Mathlib/Topology/Constructions.lean @@ -823,11 +823,14 @@ protected theorem IsOpenMap.prodMap {f : X → Y} {g : Z → W} (hf : 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) +protected theorem IsOpenEmbedding.prodMap {f : X → Y} {g : Z → W} (hf : IsOpenEmbedding f) + (hg : IsOpenEmbedding g) : IsOpenEmbedding (Prod.map f g) := + isOpenEmbedding_of_embedding_open (hf.1.prodMap hg.1) (hf.isOpenMap.prodMap hg.isOpenMap) -@[deprecated (since := "2024-10-05")] alias OpenEmbedding.prod := OpenEmbedding.prodMap +@[deprecated (since := "2024-10-18")] +alias OpenEmbedding.prodMap := IsOpenEmbedding.prodMap + +@[deprecated (since := "2024-10-05")] alias IsOpenEmbedding.prod := IsOpenEmbedding.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 @@ -902,23 +905,29 @@ theorem isOpenMap_inl : IsOpenMap (@inl X Y) := fun u hu => by theorem isOpenMap_inr : IsOpenMap (@inr X Y) := fun u hu => by simpa [isOpen_sum_iff, preimage_image_eq u Sum.inr_injective] -theorem openEmbedding_inl : OpenEmbedding (@inl X Y) := - openEmbedding_of_continuous_injective_open continuous_inl inl_injective isOpenMap_inl +theorem isOpenEmbedding_inl : IsOpenEmbedding (@inl X Y) := + isOpenEmbedding_of_continuous_injective_open continuous_inl inl_injective isOpenMap_inl + +@[deprecated (since := "2024-10-18")] +alias openEmbedding_inl := isOpenEmbedding_inl + +theorem isOpenEmbedding_inr : IsOpenEmbedding (@inr X Y) := + isOpenEmbedding_of_continuous_injective_open continuous_inr inr_injective isOpenMap_inr -theorem openEmbedding_inr : OpenEmbedding (@inr X Y) := - openEmbedding_of_continuous_injective_open continuous_inr inr_injective isOpenMap_inr +@[deprecated (since := "2024-10-18")] +alias openEmbedding_inr := isOpenEmbedding_inr theorem embedding_inl : Embedding (@inl X Y) := - openEmbedding_inl.1 + isOpenEmbedding_inl.1 theorem embedding_inr : Embedding (@inr X Y) := - openEmbedding_inr.1 + isOpenEmbedding_inr.1 theorem isOpen_range_inl : IsOpen (range (inl : X → X ⊕ Y)) := - openEmbedding_inl.2 + isOpenEmbedding_inl.2 theorem isOpen_range_inr : IsOpen (range (inr : Y → X ⊕ Y)) := - openEmbedding_inr.2 + isOpenEmbedding_inr.2 theorem isClosed_range_inl : IsClosed (range (inl : X → X ⊕ Y)) := by rw [← isOpen_compl_iff, compl_range_inl] @@ -935,10 +944,10 @@ theorem closedEmbedding_inr : ClosedEmbedding (inr : Y → X ⊕ Y) := ⟨embedding_inr, isClosed_range_inr⟩ theorem nhds_inl (x : X) : 𝓝 (inl x : X ⊕ Y) = map inl (𝓝 x) := - (openEmbedding_inl.map_nhds_eq _).symm + (isOpenEmbedding_inl.map_nhds_eq _).symm theorem nhds_inr (y : Y) : 𝓝 (inr y : X ⊕ Y) = map inr (𝓝 y) := - (openEmbedding_inr.map_nhds_eq _).symm + (isOpenEmbedding_inr.map_nhds_eq _).symm @[simp] theorem continuous_sum_map {f : X → Y} {g : Z → W} : @@ -1006,12 +1015,15 @@ theorem Continuous.subtype_val {f : Y → Subtype p} (hf : Continuous f) : Continuous fun x => (f x : X) := continuous_subtype_val.comp hf -theorem IsOpen.openEmbedding_subtype_val {s : Set X} (hs : IsOpen s) : - OpenEmbedding ((↑) : s → X) := +theorem IsOpen.isOpenEmbedding_subtypeVal {s : Set X} (hs : IsOpen s) : + IsOpenEmbedding ((↑) : s → X) := ⟨embedding_subtype_val, (@Subtype.range_coe _ s).symm ▸ hs⟩ +@[deprecated (since := "2024-10-18")] +alias IsOpen.openEmbedding_subtype_val := IsOpen.isOpenEmbedding_subtypeVal + theorem IsOpen.isOpenMap_subtype_val {s : Set X} (hs : IsOpen s) : IsOpenMap ((↑) : s → X) := - hs.openEmbedding_subtype_val.isOpenMap + hs.isOpenEmbedding_subtypeVal.isOpenMap theorem IsOpenMap.restrict {f : X → Y} (hf : IsOpenMap f) {s : Set X} (hs : IsOpen s) : IsOpenMap (s.restrict f) := @@ -1123,8 +1135,8 @@ then its restriction to the preimage of an open set is a quotient map too. -/ theorem QuotientMap.restrictPreimage_isOpen {f : X → Y} (hf : QuotientMap f) {s : Set Y} (hs : IsOpen s) : QuotientMap (s.restrictPreimage f) := by refine quotientMap_iff.2 ⟨hf.surjective.restrictPreimage _, fun U ↦ ?_⟩ - rw [hs.openEmbedding_subtype_val.open_iff_image_open, ← hf.isOpen_preimage, - (hs.preimage hf.continuous).openEmbedding_subtype_val.open_iff_image_open, + rw [hs.isOpenEmbedding_subtypeVal.open_iff_image_open, ← hf.isOpen_preimage, + (hs.preimage hf.continuous).isOpenEmbedding_subtypeVal.open_iff_image_open, image_val_preimage_restrictPreimage] open scoped Set.Notation in @@ -1528,10 +1540,13 @@ theorem isClosedMap_sigmaMk {i : ι} : IsClosedMap (@Sigma.mk ι σ i) := by theorem isClosed_range_sigmaMk {i : ι} : IsClosed (range (@Sigma.mk ι σ i)) := isClosedMap_sigmaMk.isClosed_range -theorem openEmbedding_sigmaMk {i : ι} : OpenEmbedding (@Sigma.mk ι σ i) := - openEmbedding_of_continuous_injective_open continuous_sigmaMk sigma_mk_injective +theorem isOpenEmbedding_sigmaMk {i : ι} : IsOpenEmbedding (@Sigma.mk ι σ i) := + isOpenEmbedding_of_continuous_injective_open continuous_sigmaMk sigma_mk_injective isOpenMap_sigmaMk +@[deprecated (since := "2024-10-18")] +alias openEmbedding_sigmaMk := isOpenEmbedding_sigmaMk + theorem closedEmbedding_sigmaMk {i : ι} : ClosedEmbedding (@Sigma.mk ι σ i) := closedEmbedding_of_continuous_injective_closed continuous_sigmaMk sigma_mk_injective isClosedMap_sigmaMk @@ -1540,7 +1555,7 @@ theorem embedding_sigmaMk {i : ι} : Embedding (@Sigma.mk ι σ i) := closedEmbedding_sigmaMk.1 theorem Sigma.nhds_mk (i : ι) (x : σ i) : 𝓝 (⟨i, x⟩ : Sigma σ) = Filter.map (Sigma.mk i) (𝓝 x) := - (openEmbedding_sigmaMk.map_nhds_eq x).symm + (isOpenEmbedding_sigmaMk.map_nhds_eq x).symm theorem Sigma.nhds_eq (x : Sigma σ) : 𝓝 x = Filter.map (Sigma.mk x.1) (𝓝 x.2) := by cases x @@ -1600,7 +1615,7 @@ theorem isOpenMap_sigma {f : Sigma σ → X} : IsOpenMap f ↔ ∀ i, IsOpenMap theorem isOpenMap_sigma_map {f₁ : ι → κ} {f₂ : ∀ i, σ i → τ (f₁ i)} : IsOpenMap (Sigma.map f₁ f₂) ↔ ∀ i, IsOpenMap (f₂ i) := isOpenMap_sigma.trans <| - forall_congr' fun i => (@openEmbedding_sigmaMk _ _ _ (f₁ i)).isOpenMap_iff.symm + forall_congr' fun i => (@isOpenEmbedding_sigmaMk _ _ _ (f₁ i)).isOpenMap_iff.symm theorem inducing_sigma_map {f₁ : ι → κ} {f₂ : ∀ i, σ i → τ (f₁ i)} (h₁ : Injective f₁) : Inducing (Sigma.map f₁ f₂) ↔ ∀ i, Inducing (f₂ i) := by @@ -1611,11 +1626,14 @@ theorem embedding_sigma_map {f₁ : ι → κ} {f₂ : ∀ i, σ i → τ (f₁ Embedding (Sigma.map f₁ f₂) ↔ ∀ i, Embedding (f₂ i) := by simp only [embedding_iff, Injective.sigma_map, inducing_sigma_map h, forall_and, h.sigma_map_iff] -theorem openEmbedding_sigma_map {f₁ : ι → κ} {f₂ : ∀ i, σ i → τ (f₁ i)} (h : Injective f₁) : - OpenEmbedding (Sigma.map f₁ f₂) ↔ ∀ i, OpenEmbedding (f₂ i) := by - simp only [openEmbedding_iff_embedding_open, isOpenMap_sigma_map, embedding_sigma_map h, +theorem isOpenEmbedding_sigma_map {f₁ : ι → κ} {f₂ : ∀ i, σ i → τ (f₁ i)} (h : Injective f₁) : + IsOpenEmbedding (Sigma.map f₁ f₂) ↔ ∀ i, IsOpenEmbedding (f₂ i) := by + simp only [isOpenEmbedding_iff_embedding_open, isOpenMap_sigma_map, embedding_sigma_map h, forall_and] +@[deprecated (since := "2024-10-18")] +alias openEmbedding_sigma_map := isOpenEmbedding_sigma_map + end Sigma section ULift @@ -1688,4 +1706,4 @@ theorem Filter.Eventually.prod_nhdsSet {p : X × Y → Prop} {px : X → Prop} { end NhdsSet -set_option linter.style.longFile 1700 +set_option linter.style.longFile 1900 diff --git a/Mathlib/Topology/ContinuousOn.lean b/Mathlib/Topology/ContinuousOn.lean index d69f7aec79123..ef53ffeb6a397 100644 --- a/Mathlib/Topology/ContinuousOn.lean +++ b/Mathlib/Topology/ContinuousOn.lean @@ -1081,12 +1081,15 @@ theorem Embedding.map_nhdsWithin_eq {f : α → β} (hf : Embedding f) (s : Set rw [nhdsWithin, Filter.map_inf hf.inj, hf.map_nhds_eq, map_principal, ← nhdsWithin_inter', inter_eq_self_of_subset_right (image_subset_range _ _)] -theorem OpenEmbedding.map_nhdsWithin_preimage_eq {f : α → β} (hf : OpenEmbedding f) (s : Set β) +theorem IsOpenEmbedding.map_nhdsWithin_preimage_eq {f : α → β} (hf : IsOpenEmbedding f) (s : Set β) (x : α) : map f (𝓝[f ⁻¹' s] x) = 𝓝[s] f x := by rw [hf.toEmbedding.map_nhdsWithin_eq, image_preimage_eq_inter_range] apply nhdsWithin_eq_nhdsWithin (mem_range_self _) hf.isOpen_range rw [inter_assoc, inter_self] +@[deprecated (since := "2024-10-18")] +alias OpenEmbedding.map_nhdsWithin_preimage_eq := IsOpenEmbedding.map_nhdsWithin_preimage_eq + theorem QuotientMap.continuousOn_isOpen_iff {f : α → β} {g : β → γ} (h : QuotientMap f) {s : Set β} (hs : IsOpen s) : ContinuousOn g s ↔ ContinuousOn (g ∘ f) (f ⁻¹' s) := by simp only [continuousOn_iff_continuous_restrict, (h.restrictPreimage_isOpen hs).continuous_iff] diff --git a/Mathlib/Topology/Defs/Induced.lean b/Mathlib/Topology/Defs/Induced.lean index c64958ba70443..0aaa1d46b8de9 100644 --- a/Mathlib/Topology/Defs/Induced.lean +++ b/Mathlib/Topology/Defs/Induced.lean @@ -28,7 +28,7 @@ as well as topology inducing maps, topological embeddings, and quotient maps. * `Embedding`: a map `f : X → Y` is an *embedding*, if it is a topology inducing map and it is injective. -* `OpenEmbedding`: a map `f : X → Y` is an *open embedding*, +* `IsOpenEmbedding`: a map `f : X → Y` is an *open embedding*, if it is an embedding and its range is open. An open embedding is an open map. @@ -112,10 +112,13 @@ structure Embedding [TopologicalSpace X] [TopologicalSpace Y] (f : X → Y) exte /-- An open embedding is an embedding with open range. -/ @[mk_iff] -structure OpenEmbedding (f : X → Y) extends Embedding f : Prop where +structure IsOpenEmbedding (f : X → Y) extends Embedding f : Prop where /-- The range of an open embedding is an open set. -/ isOpen_range : IsOpen <| range f +@[deprecated (since := "2024-10-18")] +alias OpenEmbedding := IsOpenEmbedding + /-- A closed embedding is an embedding with closed image. -/ @[mk_iff] structure ClosedEmbedding (f : X → Y) extends Embedding f : Prop where diff --git a/Mathlib/Topology/Gluing.lean b/Mathlib/Topology/Gluing.lean index 753550992301e..87c5a52e0a494 100644 --- a/Mathlib/Topology/Gluing.lean +++ b/Mathlib/Topology/Gluing.lean @@ -46,7 +46,7 @@ provided. * `TopCat.GlueData.preimage_range`: The preimage of the image of `U i` in `U j` is `V i j`. * `TopCat.GlueData.preimage_image_eq_image`: The preimage of the image of some `U ⊆ U i` is given by XXX. -* `TopCat.GlueData.ι_openEmbedding`: Each of the `ι i`s are open embeddings. +* `TopCat.GlueData.ι_isOpenEmbedding`: Each of the `ι i`s are open embeddings. -/ @@ -84,7 +84,7 @@ conditions are stated in a less categorical way. -/ -- porting note (#5171): removed @[nolint has_nonempty_instance] structure GlueData extends GlueData TopCat where - f_open : ∀ i j, OpenEmbedding (f i j) + f_open : ∀ i j, IsOpenEmbedding (f i j) f_mono := fun i j => (TopCat.mono_iff_injective _).mpr (f_open i j).toEmbedding.inj namespace GlueData @@ -292,10 +292,13 @@ theorem open_image_open (i : D.J) (U : Opens (𝖣.U i)) : IsOpen (𝖣.ι i '' apply (D.t j i ≫ D.f i j).continuous_toFun.isOpen_preimage exact U.isOpen -theorem ι_openEmbedding (i : D.J) : OpenEmbedding (𝖣.ι i) := - openEmbedding_of_continuous_injective_open (𝖣.ι i).continuous_toFun (D.ι_injective i) fun U h => +theorem ι_isOpenEmbedding (i : D.J) : IsOpenEmbedding (𝖣.ι i) := + isOpenEmbedding_of_continuous_injective_open (𝖣.ι i).continuous_toFun (D.ι_injective i) fun U h => D.open_image_open i ⟨U, h⟩ +@[deprecated (since := "2024-10-18")] +alias ι_openEmbedding := ι_isOpenEmbedding + /-- A family of gluing data consists of 1. An index type `J` 2. A bundled topological space `U i` for each `i : J`. @@ -356,7 +359,7 @@ def mk' (h : MkCore.{u}) : TopCat.GlueData where -- Porting note (#12129): additional beta reduction needed beta_reduce exact (h.V_id i).symm ▸ (Opens.inclusionTopIso (h.U i)).isIso_hom - f_open := fun i j : h.J => (h.V i j).openEmbedding + f_open := fun i j : h.J => (h.V i j).isOpenEmbedding t := h.t t_id i := by ext; rw [h.t_id]; rfl t' := h.t' @@ -404,7 +407,7 @@ def ofOpenSubsets : TopCat.GlueData.{u} := cocycle := fun _ _ _ _ _ => rfl } /-- The canonical map from the glue of a family of open subsets `α` into `α`. -This map is an open embedding (`fromOpenSubsetsGlue_openEmbedding`), +This map is an open embedding (`fromOpenSubsetsGlue_isOpenEmbedding`), and its range is `⋃ i, (U i : Set α)` (`range_fromOpenSubsetsGlue`). -/ def fromOpenSubsetsGlue : (ofOpenSubsets U).toGlueData.glued ⟶ TopCat.of α := @@ -439,7 +442,7 @@ theorem fromOpenSubsetsGlue_isOpenMap : IsOpenMap (fromOpenSubsetsGlue U) := by use Set.inter_subset_left constructor · rw [← Set.image_preimage_eq_inter_range] - apply (Opens.openEmbedding (X := TopCat.of α) (U i)).isOpenMap + apply (Opens.isOpenEmbedding (X := TopCat.of α) (U i)).isOpenMap convert hs i using 1 erw [← ι_fromOpenSubsetsGlue, coe_comp, Set.preimage_comp] -- porting note: `congr 1` did nothing, so I replaced it with `apply congr_arg` @@ -449,10 +452,13 @@ theorem fromOpenSubsetsGlue_isOpenMap : IsOpenMap (fromOpenSubsetsGlue U) := by rw [ι_fromOpenSubsetsGlue_apply] exact Set.mem_range_self _ -theorem fromOpenSubsetsGlue_openEmbedding : OpenEmbedding (fromOpenSubsetsGlue U) := - openEmbedding_of_continuous_injective_open (ContinuousMap.continuous_toFun _) +theorem fromOpenSubsetsGlue_isOpenEmbedding : IsOpenEmbedding (fromOpenSubsetsGlue U) := + isOpenEmbedding_of_continuous_injective_open (ContinuousMap.continuous_toFun _) (fromOpenSubsetsGlue_injective U) (fromOpenSubsetsGlue_isOpenMap U) +@[deprecated (since := "2024-10-18")] +alias fromOpenSubsetsGlue_openEmbedding := fromOpenSubsetsGlue_isOpenEmbedding + theorem range_fromOpenSubsetsGlue : Set.range (fromOpenSubsetsGlue U) = ⋃ i, (U i : Set α) := by ext constructor diff --git a/Mathlib/Topology/Homeomorph.lean b/Mathlib/Topology/Homeomorph.lean index c406c1e326c58..ce885a2ca891d 100644 --- a/Mathlib/Topology/Homeomorph.lean +++ b/Mathlib/Topology/Homeomorph.lean @@ -332,8 +332,11 @@ theorem isClosed_image (h : X ≃ₜ Y) {s : Set X} : IsClosed (h '' s) ↔ IsCl protected theorem isClosedMap (h : X ≃ₜ Y) : IsClosedMap h := fun _ => h.isClosed_image.2 -protected theorem openEmbedding (h : X ≃ₜ Y) : OpenEmbedding h := - openEmbedding_of_embedding_open h.embedding h.isOpenMap +theorem isOpenEmbedding (h : X ≃ₜ Y) : IsOpenEmbedding h := + isOpenEmbedding_of_embedding_open h.embedding h.isOpenMap + +@[deprecated (since := "2024-10-18")] +alias openEmbedding := isOpenEmbedding protected theorem closedEmbedding (h : X ≃ₜ Y) : ClosedEmbedding h := closedEmbedding_of_embedding_closed h.embedding h.isClosedMap @@ -401,7 +404,7 @@ theorem locallyConnectedSpace [i : LocallyConnectedSpace Y] (h : X ≃ₜ Y) : the domain is a locally compact space. -/ theorem locallyCompactSpace_iff (h : X ≃ₜ Y) : LocallyCompactSpace X ↔ LocallyCompactSpace Y := by - exact ⟨fun _ => h.symm.openEmbedding.locallyCompactSpace, + exact ⟨fun _ => h.symm.isOpenEmbedding.locallyCompactSpace, fun _ => h.closedEmbedding.locallyCompactSpace⟩ /-- If a bijective map `e : X ≃ Y` is continuous and open, then it is a homeomorphism. -/ @@ -908,10 +911,13 @@ protected lemma isClosedMap : IsClosedMap f := (hf.homeomorph f).isClosedMap protected lemma inducing : Inducing f := (hf.homeomorph f).inducing protected lemma quotientMap : QuotientMap f := (hf.homeomorph f).quotientMap protected lemma embedding : Embedding f := (hf.homeomorph f).embedding -protected lemma openEmbedding : OpenEmbedding f := (hf.homeomorph f).openEmbedding +lemma isOpenEmbedding : IsOpenEmbedding f := (hf.homeomorph f).isOpenEmbedding protected lemma closedEmbedding : ClosedEmbedding f := (hf.homeomorph f).closedEmbedding lemma isDenseEmbedding : IsDenseEmbedding f := (hf.homeomorph f).isDenseEmbedding +@[deprecated (since := "2024-10-18")] +alias openEmbedding := isOpenEmbedding + @[deprecated (since := "2024-09-30")] alias denseEmbedding := isDenseEmbedding @@ -932,7 +938,7 @@ lemma isHomeomorph_iff_exists_inverse : IsHomeomorph f ↔ Continuous f ∧ ∃ /-- A map is a homeomorphism iff it is a surjective embedding. -/ lemma isHomeomorph_iff_embedding_surjective : IsHomeomorph f ↔ Embedding f ∧ Surjective f where mp hf := ⟨hf.embedding, hf.surjective⟩ - mpr h := ⟨h.1.continuous, ((openEmbedding_iff f).2 ⟨h.1, h.2.range_eq ▸ isOpen_univ⟩).isOpenMap, + mpr h := ⟨h.1.continuous, ((isOpenEmbedding_iff f).2 ⟨h.1, h.2.range_eq ▸ isOpen_univ⟩).isOpenMap, h.1.inj, h.2⟩ /-- A map is a homeomorphism iff it is continuous, closed and bijective. -/ diff --git a/Mathlib/Topology/Inseparable.lean b/Mathlib/Topology/Inseparable.lean index a179072f0d072..85b288a5061a8 100644 --- a/Mathlib/Topology/Inseparable.lean +++ b/Mathlib/Topology/Inseparable.lean @@ -369,9 +369,12 @@ lemma Inducing.generalizingMap (hf : Inducing f) (h : StableUnderGeneralization obtain ⟨y, rfl⟩ := h e ⟨x, rfl⟩ exact ⟨_, hf.specializes_iff.mp e, rfl⟩ -lemma OpenEmbedding.generalizingMap (hf : OpenEmbedding f) : GeneralizingMap f := +lemma IsOpenEmbedding.generalizingMap (hf : IsOpenEmbedding f) : GeneralizingMap f := hf.toInducing.generalizingMap hf.isOpen_range.stableUnderGeneralization +@[deprecated (since := "2024-10-18")] +alias OpenEmbedding.generalizingMap := IsOpenEmbedding.generalizingMap + lemma SpecializingMap.stableUnderSpecialization_range (h : SpecializingMap f) : StableUnderSpecialization (range f) := @image_univ _ _ f ▸ stableUnderSpecialization_univ.image h diff --git a/Mathlib/Topology/Instances/ENNReal.lean b/Mathlib/Topology/Instances/ENNReal.lean index 40974db3f9717..a6719468d25d3 100644 --- a/Mathlib/Topology/Instances/ENNReal.lean +++ b/Mathlib/Topology/Instances/ENNReal.lean @@ -59,11 +59,14 @@ theorem isOpen_Ico_zero : IsOpen (Ico 0 b) := by rw [ENNReal.Ico_eq_Iio] exact isOpen_Iio -theorem openEmbedding_coe : OpenEmbedding ((↑) : ℝ≥0 → ℝ≥0∞) := +theorem isOpenEmbedding_coe : IsOpenEmbedding ((↑) : ℝ≥0 → ℝ≥0∞) := ⟨embedding_coe, by rw [range_coe']; exact isOpen_Iio⟩ +@[deprecated (since := "2024-10-18")] +alias openEmbedding_coe := isOpenEmbedding_coe + theorem coe_range_mem_nhds : range ((↑) : ℝ≥0 → ℝ≥0∞) ∈ 𝓝 (r : ℝ≥0∞) := - IsOpen.mem_nhds openEmbedding_coe.isOpen_range <| mem_range_self _ + IsOpen.mem_nhds isOpenEmbedding_coe.isOpen_range <| mem_range_self _ @[norm_cast] theorem tendsto_coe {f : Filter α} {m : α → ℝ≥0} {a : ℝ≥0} : @@ -79,7 +82,7 @@ theorem continuous_coe_iff {α} [TopologicalSpace α] {f : α → ℝ≥0} : embedding_coe.continuous_iff.symm theorem nhds_coe {r : ℝ≥0} : 𝓝 (r : ℝ≥0∞) = (𝓝 r).map (↑) := - (openEmbedding_coe.map_nhds_eq r).symm + (isOpenEmbedding_coe.map_nhds_eq r).symm theorem tendsto_nhds_coe_iff {α : Type*} {l : Filter α} {x : ℝ≥0} {f : ℝ≥0∞ → α} : Tendsto f (𝓝 ↑x) l ↔ Tendsto (f ∘ (↑) : ℝ≥0 → α) (𝓝 x) l := by @@ -91,7 +94,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.prodMap openEmbedding_coe).map_nhds_eq (r, p)).symm + ((isOpenEmbedding_coe.prodMap isOpenEmbedding_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/ENat.lean b/Mathlib/Topology/Instances/ENat.lean index ee7773df5b64b..0217f6abc247c 100644 --- a/Mathlib/Topology/Instances/ENat.lean +++ b/Mathlib/Topology/Instances/ENat.lean @@ -33,11 +33,14 @@ instance : OrderTopology ℕ∞ := ⟨rfl⟩ theorem embedding_natCast : Embedding ((↑) : ℕ → ℕ∞) := Nat.strictMono_cast.embedding_of_ordConnected <| range_natCast ▸ ordConnected_Iio -theorem openEmbedding_natCast : OpenEmbedding ((↑) : ℕ → ℕ∞) := +theorem isOpenEmbedding_natCast : IsOpenEmbedding ((↑) : ℕ → ℕ∞) := ⟨embedding_natCast, range_natCast ▸ isOpen_Iio⟩ +@[deprecated (since := "2024-10-18")] +alias openEmbedding_natCast := isOpenEmbedding_natCast + theorem nhds_natCast (n : ℕ) : 𝓝 (n : ℕ∞) = pure (n : ℕ∞) := by - simp [← openEmbedding_natCast.map_nhds_eq] + simp [← isOpenEmbedding_natCast.map_nhds_eq] @[simp] protected theorem nhds_eq_pure {n : ℕ∞} (h : n ≠ ⊤) : 𝓝 n = pure n := by diff --git a/Mathlib/Topology/Instances/EReal.lean b/Mathlib/Topology/Instances/EReal.lean index f6f81f163a7bd..1b66243074048 100644 --- a/Mathlib/Topology/Instances/EReal.lean +++ b/Mathlib/Topology/Instances/EReal.lean @@ -51,9 +51,12 @@ instance : SecondCountableTopology EReal := theorem embedding_coe : Embedding ((↑) : ℝ → EReal) := coe_strictMono.embedding_of_ordConnected <| by rw [range_coe_eq_Ioo]; exact ordConnected_Ioo -theorem openEmbedding_coe : OpenEmbedding ((↑) : ℝ → EReal) := +theorem isOpenEmbedding_coe : IsOpenEmbedding ((↑) : ℝ → EReal) := ⟨embedding_coe, by simp only [range_coe_eq_Ioo, isOpen_Ioo]⟩ +@[deprecated (since := "2024-10-18")] +alias openEmbedding_coe := isOpenEmbedding_coe + @[norm_cast] theorem tendsto_coe {α : Type*} {f : Filter α} {m : α → ℝ} {a : ℝ} : Tendsto (fun a => (m a : EReal)) f (𝓝 ↑a) ↔ Tendsto m f (𝓝 a) := @@ -66,11 +69,11 @@ theorem continuous_coe_iff {f : α → ℝ} : (Continuous fun a => (f a : EReal) embedding_coe.continuous_iff.symm theorem nhds_coe {r : ℝ} : 𝓝 (r : EReal) = (𝓝 r).map (↑) := - (openEmbedding_coe.map_nhds_eq r).symm + (isOpenEmbedding_coe.map_nhds_eq r).symm theorem nhds_coe_coe {r p : ℝ} : 𝓝 ((r : EReal), (p : EReal)) = (𝓝 (r, p)).map fun p : ℝ × ℝ => (↑p.1, ↑p.2) := - ((openEmbedding_coe.prodMap openEmbedding_coe).map_nhds_eq (r, p)).symm + ((isOpenEmbedding_coe.prodMap isOpenEmbedding_coe).map_nhds_eq (r, p)).symm theorem tendsto_toReal {a : EReal} (ha : a ≠ ⊤) (h'a : a ≠ ⊥) : Tendsto EReal.toReal (𝓝 a) (𝓝 a.toReal) := by diff --git a/Mathlib/Topology/Irreducible.lean b/Mathlib/Topology/Irreducible.lean index 17b268468df10..4855d9e0f6e57 100644 --- a/Mathlib/Topology/Irreducible.lean +++ b/Mathlib/Topology/Irreducible.lean @@ -297,7 +297,7 @@ theorem IsPreirreducible.interior (ht : IsPreirreducible t) : IsPreirreducible ( ht.open_subset isOpen_interior interior_subset theorem IsPreirreducible.preimage (ht : IsPreirreducible t) {f : Y → X} - (hf : OpenEmbedding f) : IsPreirreducible (f ⁻¹' t) := by + (hf : IsOpenEmbedding f) : IsPreirreducible (f ⁻¹' t) := by rintro U V hU hV ⟨x, hx, hx'⟩ ⟨y, hy, hy'⟩ obtain ⟨_, h₁, ⟨y, h₂, rfl⟩, ⟨y', h₃, h₄⟩⟩ := ht _ _ (hf.isOpenMap _ hU) (hf.isOpenMap _ hV) ⟨f x, hx, Set.mem_image_of_mem f hx'⟩ diff --git a/Mathlib/Topology/IsLocalHomeomorph.lean b/Mathlib/Topology/IsLocalHomeomorph.lean index 6351032d8967e..d75d32d051089 100644 --- a/Mathlib/Topology/IsLocalHomeomorph.lean +++ b/Mathlib/Topology/IsLocalHomeomorph.lean @@ -39,22 +39,26 @@ the source of some `e : PartialHomeomorph X Y` with `f = e`. -/ def IsLocalHomeomorphOn := ∀ x ∈ s, ∃ e : PartialHomeomorph X Y, x ∈ e.source ∧ f = e -theorem isLocalHomeomorphOn_iff_openEmbedding_restrict {f : X → Y} : - IsLocalHomeomorphOn f s ↔ ∀ x ∈ s, ∃ U ∈ 𝓝 x, OpenEmbedding (U.restrict f) := by +theorem isLocalHomeomorphOn_iff_isOpenEmbedding_restrict {f : X → Y} : + IsLocalHomeomorphOn f s ↔ ∀ x ∈ s, ∃ U ∈ 𝓝 x, IsOpenEmbedding (U.restrict f) := by refine ⟨fun h x hx ↦ ?_, fun h x hx ↦ ?_⟩ · obtain ⟨e, hxe, rfl⟩ := h x hx - exact ⟨e.source, e.open_source.mem_nhds hxe, e.openEmbedding_restrict⟩ + exact ⟨e.source, e.open_source.mem_nhds hxe, e.isOpenEmbedding_restrict⟩ · obtain ⟨U, hU, emb⟩ := h x hx - have : OpenEmbedding ((interior U).restrict f) := by + have : IsOpenEmbedding ((interior U).restrict f) := by refine emb.comp ⟨embedding_inclusion interior_subset, ?_⟩ rw [Set.range_inclusion]; exact isOpen_induced isOpen_interior - obtain ⟨cont, inj, openMap⟩ := openEmbedding_iff_continuous_injective_open.mp this + obtain ⟨cont, inj, openMap⟩ := isOpenEmbedding_iff_continuous_injective_open.mp this haveI : Nonempty X := ⟨x⟩ exact ⟨PartialHomeomorph.ofContinuousOpenRestrict (Set.injOn_iff_injective.mpr inj).toPartialEquiv (continuousOn_iff_continuous_restrict.mpr cont) openMap isOpen_interior, mem_interior_iff_mem_nhds.mpr hU, rfl⟩ +@[deprecated (since := "2024-10-18")] +alias isLocalHomeomorphOn_iff_openEmbedding_restrict := + isLocalHomeomorphOn_iff_isOpenEmbedding_restrict + namespace IsLocalHomeomorphOn /-- Proves that `f` satisfies `IsLocalHomeomorphOn f s`. The condition `h` is weaker than the @@ -141,14 +145,20 @@ theorem isLocalHomeomorph_iff_isLocalHomeomorphOn_univ : protected theorem IsLocalHomeomorph.isLocalHomeomorphOn (hf : IsLocalHomeomorph f) : IsLocalHomeomorphOn f s := fun x _ ↦ hf x -theorem isLocalHomeomorph_iff_openEmbedding_restrict {f : X → Y} : - IsLocalHomeomorph f ↔ ∀ x : X, ∃ U ∈ 𝓝 x, OpenEmbedding (U.restrict f) := by +theorem isLocalHomeomorph_iff_isOpenEmbedding_restrict {f : X → Y} : + IsLocalHomeomorph f ↔ ∀ x : X, ∃ U ∈ 𝓝 x, IsOpenEmbedding (U.restrict f) := by simp_rw [isLocalHomeomorph_iff_isLocalHomeomorphOn_univ, - isLocalHomeomorphOn_iff_openEmbedding_restrict, imp_iff_right (Set.mem_univ _)] + isLocalHomeomorphOn_iff_isOpenEmbedding_restrict, imp_iff_right (Set.mem_univ _)] + +@[deprecated (since := "2024-10-18")] +alias isLocalHomeomorph_iff_openEmbedding_restrict := isLocalHomeomorph_iff_isOpenEmbedding_restrict + +theorem IsOpenEmbedding.isLocalHomeomorph (hf : IsOpenEmbedding f) : IsLocalHomeomorph f := + isLocalHomeomorph_iff_isOpenEmbedding_restrict.mpr fun _ ↦ + ⟨_, Filter.univ_mem, hf.comp (Homeomorph.Set.univ X).isOpenEmbedding⟩ -theorem OpenEmbedding.isLocalHomeomorph (hf : OpenEmbedding f) : IsLocalHomeomorph f := - isLocalHomeomorph_iff_openEmbedding_restrict.mpr fun _ ↦ - ⟨_, Filter.univ_mem, hf.comp (Homeomorph.Set.univ X).openEmbedding⟩ +@[deprecated (since := "2024-10-18")] +alias OpenEmbedding.isLocalHomeomorph := IsOpenEmbedding.isLocalHomeomorph variable (f) @@ -194,15 +204,18 @@ protected theorem comp (hg : IsLocalHomeomorph g) (hf : IsLocalHomeomorph f) : (hg.isLocalHomeomorphOn.comp hf.isLocalHomeomorphOn (Set.univ.mapsTo_univ f)) /-- An injective local homeomorphism is an open embedding. -/ -theorem openEmbedding_of_injective (hf : IsLocalHomeomorph f) (hi : f.Injective) : - OpenEmbedding f := - openEmbedding_of_continuous_injective_open hf.continuous hi hf.isOpenMap +theorem isOpenEmbedding_of_injective (hf : IsLocalHomeomorph f) (hi : f.Injective) : + IsOpenEmbedding f := + isOpenEmbedding_of_continuous_injective_open hf.continuous hi hf.isOpenMap + +@[deprecated (since := "2024-10-18")] +alias openEmbedding_of_injective := isOpenEmbedding_of_injective /-- A surjective embedding is a homeomorphism. -/ noncomputable def _root_.Embedding.toHomeomeomorph_of_surjective (hf : Embedding f) (hsurj : Function.Surjective f) : X ≃ₜ Y := Homeomorph.homeomorphOfContinuousOpen (Equiv.ofBijective f ⟨hf.inj, hsurj⟩) - hf.continuous (hf.toOpenEmbedding_of_surjective hsurj).isOpenMap + hf.continuous (hf.toIsOpenEmbedding_of_surjective hsurj).isOpenMap /-- A bijective local homeomorphism is a homeomorphism. -/ noncomputable def toHomeomorph_of_bijective (hf : IsLocalHomeomorph f) (hb : f.Bijective) : @@ -210,9 +223,12 @@ noncomputable def toHomeomorph_of_bijective (hf : IsLocalHomeomorph f) (hb : f.B Homeomorph.homeomorphOfContinuousOpen (Equiv.ofBijective f hb) hf.continuous hf.isOpenMap /-- Continuous local sections of a local homeomorphism are open embeddings. -/ -theorem openEmbedding_of_comp (hf : IsLocalHomeomorph g) (hgf : OpenEmbedding (g ∘ f)) - (cont : Continuous f) : OpenEmbedding f := - (hgf.isLocalHomeomorph.of_comp hf cont).openEmbedding_of_injective hgf.inj.of_comp +theorem isOpenEmbedding_of_comp (hf : IsLocalHomeomorph g) (hgf : IsOpenEmbedding (g ∘ f)) + (cont : Continuous f) : IsOpenEmbedding f := + (hgf.isLocalHomeomorph.of_comp hf cont).isOpenEmbedding_of_injective hgf.inj.of_comp + +@[deprecated (since := "2024-10-18")] +alias openEmbedding_of_comp := isOpenEmbedding_of_comp open TopologicalSpace in /-- Ranges of continuous local sections of a local homeomorphism @@ -221,7 +237,7 @@ theorem isTopologicalBasis (hf : IsLocalHomeomorph f) : IsTopologicalBasis {U : Set X | ∃ V : Set Y, IsOpen V ∧ ∃ s : C(V,X), f ∘ s = (↑) ∧ Set.range s = U} := by refine isTopologicalBasis_of_isOpen_of_nhds ?_ fun x U hx hU ↦ ?_ · rintro _ ⟨U, hU, s, hs, rfl⟩ - refine (openEmbedding_of_comp hf (hs ▸ ⟨embedding_subtype_val, ?_⟩) s.continuous).isOpen_range + refine (isOpenEmbedding_of_comp hf (hs ▸ ⟨embedding_subtype_val, ?_⟩) s.continuous).isOpen_range rwa [Subtype.range_val] · obtain ⟨f, hxf, rfl⟩ := hf x refine ⟨f.source ∩ U, ⟨f.target ∩ f.symm ⁻¹' U, f.symm.isOpen_inter_preimage hU, diff --git a/Mathlib/Topology/LocalAtTarget.lean b/Mathlib/Topology/LocalAtTarget.lean index 184b90a9c30e4..4e1a6f06f340a 100644 --- a/Mathlib/Topology/LocalAtTarget.lean +++ b/Mathlib/Topology/LocalAtTarget.lean @@ -12,7 +12,7 @@ import Mathlib.Topology.LocallyClosed We show that the following properties of continuous maps are local at the target : - `Inducing` - `Embedding` -- `OpenEmbedding` +- `IsOpenEmbedding` - `ClosedEmbedding` -/ @@ -40,12 +40,18 @@ theorem Set.restrictPreimage_embedding (s : Set β) (h : Embedding f) : alias Embedding.restrictPreimage := Set.restrictPreimage_embedding -theorem Set.restrictPreimage_openEmbedding (s : Set β) (h : OpenEmbedding f) : - OpenEmbedding (s.restrictPreimage f) := +theorem Set.restrictPreimage_isOpenEmbedding (s : Set β) (h : IsOpenEmbedding f) : + IsOpenEmbedding (s.restrictPreimage f) := ⟨h.1.restrictPreimage s, (s.range_restrictPreimage f).symm ▸ continuous_subtype_val.isOpen_preimage _ h.isOpen_range⟩ -alias OpenEmbedding.restrictPreimage := Set.restrictPreimage_openEmbedding +@[deprecated (since := "2024-10-18")] +alias Set.restrictPreimage_openEmbedding := Set.restrictPreimage_isOpenEmbedding + +alias IsOpenEmbedding.restrictPreimage := Set.restrictPreimage_isOpenEmbedding + +@[deprecated (since := "2024-10-18")] +alias OpenEmbedding.restrictPreimage := IsOpenEmbedding.restrictPreimage theorem Set.restrictPreimage_closedEmbedding (s : Set β) (h : ClosedEmbedding f) : ClosedEmbedding (s.restrictPreimage f) := @@ -96,7 +102,7 @@ theorem isOpen_iff_coe_preimage_of_iSup_eq_top (s : Set β) : -- Porting note: rewrote to avoid ´simp´ issues rw [isOpen_iff_inter_of_iSup_eq_top hU s] refine forall_congr' fun i => ?_ - rw [(U _).2.openEmbedding_subtype_val.open_iff_image_open] + rw [(U _).2.isOpenEmbedding_subtypeVal.open_iff_image_open] erw [Set.image_preimage_eq_inter_range] rw [Subtype.range_coe, Opens.carrier_eq_coe] @@ -110,7 +116,7 @@ theorem isLocallyClosed_iff_coe_preimage_of_iSup_eq_top (s : Set β) : rw [isOpen_iff_coe_preimage_of_iSup_eq_top hU] exact forall_congr' fun i ↦ by have : coborder ((↑) ⁻¹' s : Set (U i)) = Subtype.val ⁻¹' coborder s := by - exact (U i).isOpen.openEmbedding_subtype_val.coborder_preimage _ + exact (U i).isOpen.isOpenEmbedding_subtypeVal.coborder_preimage _ rw [this] theorem isOpenMap_iff_isOpenMap_of_iSup_eq_top : @@ -150,7 +156,7 @@ theorem inducing_iff_inducing_of_iSup_eq_top (h : Continuous f) : (show f x ∈ iSup U by rw [hU] trivial) - rw [← OpenEmbedding.map_nhds_eq (h.1 _ (U i).2).openEmbedding_subtype_val ⟨x, hi⟩, + rw [← IsOpenEmbedding.map_nhds_eq (h.1 _ (U i).2).isOpenEmbedding_subtypeVal ⟨x, hi⟩, (H i) ⟨x, hi⟩, Filter.subtype_coe_map_comap, Function.comp_apply, Subtype.coe_mk, inf_eq_left, Filter.le_principal_iff] exact Filter.preimage_mem_comap ((U i).2.mem_nhds hi) @@ -165,15 +171,19 @@ theorem embedding_iff_embedding_of_iSup_eq_top (h : Continuous f) : convert congr_arg SetLike.coe hU simp -theorem openEmbedding_iff_openEmbedding_of_iSup_eq_top (h : Continuous f) : - OpenEmbedding f ↔ ∀ i, OpenEmbedding ((U i).1.restrictPreimage f) := by - simp_rw [openEmbedding_iff] +theorem isOpenEmbedding_iff_isOpenEmbedding_of_iSup_eq_top (h : Continuous f) : + IsOpenEmbedding f ↔ ∀ i, IsOpenEmbedding ((U i).1.restrictPreimage f) := by + simp_rw [isOpenEmbedding_iff] rw [forall_and] apply and_congr · apply embedding_iff_embedding_of_iSup_eq_top <;> assumption · simp_rw [Set.range_restrictPreimage] apply isOpen_iff_coe_preimage_of_iSup_eq_top hU +@[deprecated (since := "2024-10-18")] +alias openEmbedding_iff_openEmbedding_of_iSup_eq_top := + isOpenEmbedding_iff_isOpenEmbedding_of_iSup_eq_top + theorem closedEmbedding_iff_closedEmbedding_of_iSup_eq_top (h : Continuous f) : ClosedEmbedding f ↔ ∀ i, ClosedEmbedding ((U i).1.restrictPreimage f) := by simp_rw [closedEmbedding_iff] diff --git a/Mathlib/Topology/LocallyClosed.lean b/Mathlib/Topology/LocallyClosed.lean index fd260f0e11c03..6aa387864372e 100644 --- a/Mathlib/Topology/LocallyClosed.lean +++ b/Mathlib/Topology/LocallyClosed.lean @@ -91,10 +91,13 @@ lemma coborder_preimage (hf : IsOpenMap f) (hf' : Continuous f) (s : Set Y) : (hf.coborder_preimage_subset s).antisymm (hf'.preimage_coborder_subset s) protected -lemma OpenEmbedding.coborder_preimage (hf : OpenEmbedding f) (s : Set Y) : +lemma IsOpenEmbedding.coborder_preimage (hf : IsOpenEmbedding f) (s : Set Y) : coborder (f ⁻¹' s) = f ⁻¹' (coborder s) := coborder_preimage hf.isOpenMap hf.continuous s +@[deprecated (since := "2024-10-18")] +alias OpenEmbedding.coborder_preimage := IsOpenEmbedding.coborder_preimage + lemma isClosed_preimage_val_coborder : IsClosed (coborder s ↓∩ s) := by rw [isClosed_preimage_val, inter_eq_right.mpr subset_coborder, coborder_inter_closure] diff --git a/Mathlib/Topology/Maps/Basic.lean b/Mathlib/Topology/Maps/Basic.lean index de7a4ffef4499..c4aaea387693f 100644 --- a/Mathlib/Topology/Maps/Basic.lean +++ b/Mathlib/Topology/Maps/Basic.lean @@ -20,7 +20,7 @@ This file introduces the following properties of a map `f : X → Y` between top are identified by `f` are also inseparable in the topology on `X`. * `Embedding f` means `f` is inducing and also injective. Equivalently, `f` identifies `X` with a subspace of `Y`. -* `OpenEmbedding f` means `f` is an embedding with open image, so it identifies `X` with an +* `IsOpenEmbedding f` means `f` is an embedding with open image, so it identifies `X` with an open subspace of `Y`. Equivalently, `f` is an embedding and an open map. * `ClosedEmbedding f` similarly means `f` is an embedding with closed image, so it identifies `X` with a closed subspace of `Y`. Equivalently, `f` is an embedding and a closed map. @@ -470,99 +470,138 @@ theorem IsClosedMap.mapClusterPt_iff_lift'_closure end IsClosedMap -section OpenEmbedding +section IsOpenEmbedding variable [TopologicalSpace X] [TopologicalSpace Y] -theorem OpenEmbedding.isOpenMap (hf : OpenEmbedding f) : IsOpenMap f := +theorem IsOpenEmbedding.isOpenMap (hf : IsOpenEmbedding f) : IsOpenMap f := hf.toEmbedding.toInducing.isOpenMap hf.isOpen_range -theorem OpenEmbedding.map_nhds_eq (hf : OpenEmbedding f) (x : X) : +@[deprecated (since := "2024-10-18")] +alias OpenEmbedding.isOpenMap := IsOpenEmbedding.isOpenMap + +theorem IsOpenEmbedding.map_nhds_eq (hf : IsOpenEmbedding f) (x : X) : map f (𝓝 x) = 𝓝 (f x) := hf.toEmbedding.map_nhds_of_mem _ <| hf.isOpen_range.mem_nhds <| mem_range_self _ -theorem OpenEmbedding.open_iff_image_open (hf : OpenEmbedding f) {s : Set X} : +@[deprecated (since := "2024-10-18")] +alias OpenEmbedding.map_nhds_eq := IsOpenEmbedding.map_nhds_eq + +theorem IsOpenEmbedding.open_iff_image_open (hf : IsOpenEmbedding f) {s : Set X} : IsOpen s ↔ IsOpen (f '' s) := ⟨hf.isOpenMap s, fun h => by convert ← h.preimage hf.toEmbedding.continuous apply preimage_image_eq _ hf.inj⟩ -theorem OpenEmbedding.tendsto_nhds_iff [TopologicalSpace Z] {f : ι → Y} {l : Filter ι} {y : Y} - (hg : OpenEmbedding g) : Tendsto f l (𝓝 y) ↔ Tendsto (g ∘ f) l (𝓝 (g y)) := +@[deprecated (since := "2024-10-18")] +alias OpenEmbedding.open_iff_image_open := IsOpenEmbedding.open_iff_image_open + +theorem IsOpenEmbedding.tendsto_nhds_iff [TopologicalSpace Z] {f : ι → Y} {l : Filter ι} {y : Y} + (hg : IsOpenEmbedding g) : Tendsto f l (𝓝 y) ↔ Tendsto (g ∘ f) l (𝓝 (g y)) := hg.toEmbedding.tendsto_nhds_iff -theorem OpenEmbedding.tendsto_nhds_iff' (hf : OpenEmbedding f) {l : Filter Z} {x : X} : +@[deprecated (since := "2024-10-18")] +alias OpenEmbedding.tendsto_nhds_iff := IsOpenEmbedding.tendsto_nhds_iff + +theorem IsOpenEmbedding.tendsto_nhds_iff' (hf : IsOpenEmbedding f) {l : Filter Z} {x : X} : Tendsto (g ∘ f) (𝓝 x) l ↔ Tendsto g (𝓝 (f x)) l := by rw [Tendsto, ← map_map, hf.map_nhds_eq]; rfl -theorem OpenEmbedding.continuousAt_iff [TopologicalSpace Z] (hf : OpenEmbedding f) {x : X} : +@[deprecated (since := "2024-10-18")] +alias OpenEmbedding.tendsto_nhds_iff' := IsOpenEmbedding.tendsto_nhds_iff' + +theorem IsOpenEmbedding.continuousAt_iff [TopologicalSpace Z] (hf : IsOpenEmbedding f) {x : X} : ContinuousAt (g ∘ f) x ↔ ContinuousAt g (f x) := hf.tendsto_nhds_iff' -theorem OpenEmbedding.continuous (hf : OpenEmbedding f) : Continuous f := +@[deprecated (since := "2024-10-18")] +alias OpenEmbedding.continuousAt_iff := IsOpenEmbedding.continuousAt_iff + +theorem IsOpenEmbedding.continuous (hf : IsOpenEmbedding f) : Continuous f := hf.toEmbedding.continuous -theorem OpenEmbedding.open_iff_preimage_open (hf : OpenEmbedding f) {s : Set Y} +@[deprecated (since := "2024-10-18")] +alias OpenEmbedding.continuous := IsOpenEmbedding.continuous + +theorem IsOpenEmbedding.open_iff_preimage_open (hf : IsOpenEmbedding f) {s : Set Y} (hs : s ⊆ range f) : IsOpen s ↔ IsOpen (f ⁻¹' s) := by rw [hf.open_iff_image_open, image_preimage_eq_inter_range, inter_eq_self_of_subset_left hs] -theorem openEmbedding_of_embedding_open (h₁ : Embedding f) (h₂ : IsOpenMap f) : - OpenEmbedding f := +@[deprecated (since := "2024-10-18")] +alias OpenEmbedding.open_iff_preimage_open := IsOpenEmbedding.open_iff_preimage_open + +theorem isOpenEmbedding_of_embedding_open (h₁ : Embedding f) (h₂ : IsOpenMap f) : + IsOpenEmbedding f := ⟨h₁, h₂.isOpen_range⟩ -/-- A surjective embedding is an `OpenEmbedding`. -/ -theorem _root_.Embedding.toOpenEmbedding_of_surjective (hf : Embedding f) (hsurj : f.Surjective) : - OpenEmbedding f := +@[deprecated (since := "2024-10-18")] +alias openEmbedding_of_embedding_open := isOpenEmbedding_of_embedding_open + +/-- A surjective embedding is an `IsOpenEmbedding`. -/ +theorem _root_.Embedding.toIsOpenEmbedding_of_surjective (hf : Embedding f) (hsurj : f.Surjective) : + IsOpenEmbedding f := ⟨hf, hsurj.range_eq ▸ isOpen_univ⟩ -theorem openEmbedding_iff_embedding_open : - OpenEmbedding f ↔ Embedding f ∧ IsOpenMap f := - ⟨fun h => ⟨h.1, h.isOpenMap⟩, fun h => openEmbedding_of_embedding_open h.1 h.2⟩ +@[deprecated (since := "2024-10-18")] +alias _root_.Embedding.toOpenEmbedding_of_surjective := Embedding.toIsOpenEmbedding_of_surjective + +theorem isOpenEmbedding_iff_embedding_open : + IsOpenEmbedding f ↔ Embedding f ∧ IsOpenMap f := + ⟨fun h => ⟨h.1, h.isOpenMap⟩, fun h => isOpenEmbedding_of_embedding_open h.1 h.2⟩ -theorem openEmbedding_of_continuous_injective_open - (h₁ : Continuous f) (h₂ : Injective f) (h₃ : IsOpenMap f) : OpenEmbedding f := by - simp only [openEmbedding_iff_embedding_open, embedding_iff, inducing_iff_nhds, *, and_true] +@[deprecated (since := "2024-10-18")] +alias openEmbedding_iff_embedding_open := isOpenEmbedding_iff_embedding_open + +theorem isOpenEmbedding_of_continuous_injective_open + (h₁ : Continuous f) (h₂ : Injective f) (h₃ : IsOpenMap f) : IsOpenEmbedding f := by + simp only [isOpenEmbedding_iff_embedding_open, embedding_iff, inducing_iff_nhds, *, and_true] exact fun x => le_antisymm (h₁.tendsto _).le_comap (@comap_map _ _ (𝓝 x) _ h₂ ▸ comap_mono (h₃.nhds_le _)) -theorem openEmbedding_iff_continuous_injective_open : - OpenEmbedding f ↔ Continuous f ∧ Injective f ∧ IsOpenMap f := +theorem isOpenEmbedding_iff_continuous_injective_open : + IsOpenEmbedding f ↔ Continuous f ∧ Injective f ∧ IsOpenMap f := ⟨fun h => ⟨h.continuous, h.inj, h.isOpenMap⟩, fun h => - openEmbedding_of_continuous_injective_open h.1 h.2.1 h.2.2⟩ + isOpenEmbedding_of_continuous_injective_open h.1 h.2.1 h.2.2⟩ + +@[deprecated (since := "2024-10-18")] +alias openEmbedding_iff_continuous_injective_open := isOpenEmbedding_iff_continuous_injective_open -theorem openEmbedding_id : OpenEmbedding (@id X) := +theorem isOpenEmbedding_id : IsOpenEmbedding (@id X) := ⟨embedding_id, IsOpenMap.id.isOpen_range⟩ -namespace OpenEmbedding +@[deprecated (since := "2024-10-18")] +alias openEmbedding_id := isOpenEmbedding_id + +namespace IsOpenEmbedding variable [TopologicalSpace Z] -protected theorem comp (hg : OpenEmbedding g) - (hf : OpenEmbedding f) : OpenEmbedding (g ∘ f) := +protected theorem comp (hg : IsOpenEmbedding g) + (hf : IsOpenEmbedding f) : IsOpenEmbedding (g ∘ f) := ⟨hg.1.comp hf.1, (hg.isOpenMap.comp hf.isOpenMap).isOpen_range⟩ -theorem isOpenMap_iff (hg : OpenEmbedding g) : +theorem isOpenMap_iff (hg : IsOpenEmbedding g) : IsOpenMap f ↔ IsOpenMap (g ∘ f) := by simp_rw [isOpenMap_iff_nhds_le, ← map_map, comp, ← hg.map_nhds_eq, Filter.map_le_map_iff hg.inj] -theorem of_comp_iff (f : X → Y) (hg : OpenEmbedding g) : - OpenEmbedding (g ∘ f) ↔ OpenEmbedding f := by - simp only [openEmbedding_iff_continuous_injective_open, ← hg.isOpenMap_iff, ← +theorem of_comp_iff (f : X → Y) (hg : IsOpenEmbedding g) : + IsOpenEmbedding (g ∘ f) ↔ IsOpenEmbedding f := by + simp only [isOpenEmbedding_iff_continuous_injective_open, ← hg.isOpenMap_iff, ← hg.1.continuous_iff, hg.inj.of_comp_iff] -theorem of_comp (f : X → Y) (hg : OpenEmbedding g) - (h : OpenEmbedding (g ∘ f)) : OpenEmbedding f := - (OpenEmbedding.of_comp_iff f hg).1 h +theorem of_comp (f : X → Y) (hg : IsOpenEmbedding g) + (h : IsOpenEmbedding (g ∘ f)) : IsOpenEmbedding f := + (IsOpenEmbedding.of_comp_iff f hg).1 h -theorem of_isEmpty [IsEmpty X] (f : X → Y) : OpenEmbedding f := - openEmbedding_of_embedding_open (.of_subsingleton f) (IsOpenMap.of_isEmpty f) +theorem of_isEmpty [IsEmpty X] (f : X → Y) : IsOpenEmbedding f := + isOpenEmbedding_of_embedding_open (.of_subsingleton f) (IsOpenMap.of_isEmpty f) -theorem image_mem_nhds {f : X → Y} (hf : OpenEmbedding f) {s : Set X} {x : X} : +theorem image_mem_nhds {f : X → Y} (hf : IsOpenEmbedding f) {s : Set X} {x : X} : f '' s ∈ 𝓝 (f x) ↔ s ∈ 𝓝 x := by rw [← hf.map_nhds_eq, mem_map, preimage_image_eq _ hf.inj] -end OpenEmbedding +end IsOpenEmbedding -end OpenEmbedding +end IsOpenEmbedding section ClosedEmbedding diff --git a/Mathlib/Topology/PartialHomeomorph.lean b/Mathlib/Topology/PartialHomeomorph.lean index 3cecb43406b4e..4d2b050b3eb3c 100644 --- a/Mathlib/Topology/PartialHomeomorph.lean +++ b/Mathlib/Topology/PartialHomeomorph.lean @@ -1092,18 +1092,24 @@ def toHomeomorphOfSourceEqUnivTargetEqUniv (h : e.source = (univ : Set X)) (h' : continuous_invFun := by simpa only [continuous_iff_continuousOn_univ, h'] using e.continuousOn_symm -theorem openEmbedding_restrict : OpenEmbedding (e.source.restrict e) := by - refine openEmbedding_of_continuous_injective_open (e.continuousOn.comp_continuous +theorem isOpenEmbedding_restrict : IsOpenEmbedding (e.source.restrict e) := by + refine isOpenEmbedding_of_continuous_injective_open (e.continuousOn.comp_continuous continuous_subtype_val Subtype.prop) e.injOn.injective fun V hV ↦ ?_ rw [Set.restrict_eq, Set.image_comp] exact e.isOpen_image_of_subset_source (e.open_source.isOpenMap_subtype_val V hV) fun _ ⟨x, _, h⟩ ↦ h ▸ x.2 +@[deprecated (since := "2024-10-18")] +alias openEmbedding_restrict := isOpenEmbedding_restrict + /-- A partial homeomorphism whose source is all of `X` defines an open embedding of `X` into `Y`. -The converse is also true; see `OpenEmbedding.toPartialHomeomorph`. -/ -theorem to_openEmbedding (h : e.source = Set.univ) : OpenEmbedding e := - e.openEmbedding_restrict.comp - ((Homeomorph.setCongr h).trans <| Homeomorph.Set.univ X).symm.openEmbedding +The converse is also true; see `IsOpenEmbedding.toPartialHomeomorph`. -/ +theorem to_isOpenEmbedding (h : e.source = Set.univ) : IsOpenEmbedding e := + e.isOpenEmbedding_restrict.comp + ((Homeomorph.setCongr h).trans <| Homeomorph.Set.univ X).symm.isOpenEmbedding + +@[deprecated (since := "2024-10-18")] +alias to_openEmbedding := to_isOpenEmbedding end PartialHomeomorph @@ -1156,12 +1162,13 @@ theorem trans_transPartialHomeomorph (e : X ≃ₜ Y) (e' : Y ≃ₜ Z) (f'' : P end Homeomorph -namespace OpenEmbedding +namespace IsOpenEmbedding -variable (f : X → Y) (h : OpenEmbedding f) +variable (f : X → Y) (h : IsOpenEmbedding f) /-- An open embedding of `X` into `Y`, with `X` nonempty, defines a partial homeomorphism -whose source is all of `X`. The converse is also true; see `PartialHomeomorph.to_openEmbedding`. -/ +whose source is all of `X`. The converse is also true; see +`PartialHomeomorph.to_isOpenEmbedding`. -/ @[simps! (config := mfld_cfg) apply source target] noncomputable def toPartialHomeomorph [Nonempty X] : PartialHomeomorph X Y := PartialHomeomorph.ofContinuousOpen (h.toEmbedding.inj.injOn.toPartialEquiv f univ) @@ -1178,7 +1185,7 @@ lemma toPartialHomeomorph_right_inv {x : Y} (hx : x ∈ Set.range f) : rw [← congr_fun (h.toPartialHomeomorph_apply f), PartialHomeomorph.right_inv] rwa [toPartialHomeomorph_target] -end OpenEmbedding +end IsOpenEmbedding /-! inclusion of an open set in a topological space -/ namespace TopologicalSpace.Opens @@ -1191,7 +1198,7 @@ variable (s : Opens X) (hs : Nonempty s) /-- The inclusion of an open subset `s` of a space `X` into `X` is a partial homeomorphism from the subtype `s` to `X`. -/ noncomputable def partialHomeomorphSubtypeCoe : PartialHomeomorph s X := - OpenEmbedding.toPartialHomeomorph _ s.2.openEmbedding_subtype_val + IsOpenEmbedding.toPartialHomeomorph _ s.2.isOpenEmbedding_subtypeVal @[simp, mfld_simps] theorem partialHomeomorphSubtypeCoe_coe : (s.partialHomeomorphSubtypeCoe hs : s → X) = (↑) := diff --git a/Mathlib/Topology/QuasiSeparated.lean b/Mathlib/Topology/QuasiSeparated.lean index 869b89e96ebba..73240b9ea2c7e 100644 --- a/Mathlib/Topology/QuasiSeparated.lean +++ b/Mathlib/Topology/QuasiSeparated.lean @@ -22,7 +22,7 @@ open subsets, but their intersection `(0, 1]` is not. of any pairs of compact open subsets of `s` are still compact. - `QuasiSeparatedSpace`: A topological space is quasi-separated if the intersections of any pairs of compact open subsets are still compact. -- `QuasiSeparatedSpace.of_openEmbedding`: If `f : α → β` is an open embedding, and `β` is +- `QuasiSeparatedSpace.of_isOpenEmbedding`: If `f : α → β` is an open embedding, and `β` is a quasi-separated space, then so is `α`. -/ @@ -79,7 +79,7 @@ theorem IsQuasiSeparated.image_of_embedding {s : Set α} (H : IsQuasiSeparated s rw [Set.image_preimage_eq_inter_range, Set.inter_eq_left] exact hV.trans (Set.image_subset_range _ _) -theorem OpenEmbedding.isQuasiSeparated_iff (h : OpenEmbedding f) {s : Set α} : +theorem IsOpenEmbedding.isQuasiSeparated_iff (h : IsOpenEmbedding f) {s : Set α} : IsQuasiSeparated s ↔ IsQuasiSeparated (f '' s) := by refine ⟨fun hs => hs.image_of_embedding h.toEmbedding, ?_⟩ intro H U V hU hU' hU'' hV hV' hV'' @@ -88,10 +88,13 @@ theorem OpenEmbedding.isQuasiSeparated_iff (h : OpenEmbedding f) {s : Set α} : H (f '' U) (f '' V) (Set.image_subset _ hU) (h.isOpenMap _ hU') (hU''.image h.continuous) (Set.image_subset _ hV) (h.isOpenMap _ hV') (hV''.image h.continuous) +@[deprecated (since := "2024-10-18")] +alias OpenEmbedding.isQuasiSeparated_iff := IsOpenEmbedding.isQuasiSeparated_iff + theorem isQuasiSeparated_iff_quasiSeparatedSpace (s : Set α) (hs : IsOpen s) : IsQuasiSeparated s ↔ QuasiSeparatedSpace s := by rw [← isQuasiSeparated_univ_iff] - convert (hs.openEmbedding_subtype_val.isQuasiSeparated_iff (s := Set.univ)).symm + convert (hs.isOpenEmbedding_subtypeVal.isQuasiSeparated_iff (s := Set.univ)).symm simp theorem IsQuasiSeparated.of_subset {s t : Set α} (ht : IsQuasiSeparated t) (h : s ⊆ t) : @@ -110,7 +113,10 @@ theorem IsQuasiSeparated.of_quasiSeparatedSpace (s : Set α) [QuasiSeparatedSpac IsQuasiSeparated s := isQuasiSeparated_univ.of_subset (Set.subset_univ _) -theorem QuasiSeparatedSpace.of_openEmbedding (h : OpenEmbedding f) [QuasiSeparatedSpace β] : +theorem QuasiSeparatedSpace.of_isOpenEmbedding (h : IsOpenEmbedding f) [QuasiSeparatedSpace β] : QuasiSeparatedSpace α := isQuasiSeparated_univ_iff.mp (h.isQuasiSeparated_iff.mpr <| IsQuasiSeparated.of_quasiSeparatedSpace _) + +@[deprecated (since := "2024-10-18")] +alias QuasiSeparatedSpace.of_openEmbedding := QuasiSeparatedSpace.of_isOpenEmbedding diff --git a/Mathlib/Topology/SeparatedMap.lean b/Mathlib/Topology/SeparatedMap.lean index 471b35d8086e1..ccb0ea0f688df 100644 --- a/Mathlib/Topology/SeparatedMap.lean +++ b/Mathlib/Topology/SeparatedMap.lean @@ -141,15 +141,18 @@ theorem isLocallyInjective_iff_isOpen_diagonal {f : X → Y} : exact ⟨t₁ ∩ t₂, Filter.inter_mem h₁ h₂, fun x₁ h₁ x₂ h₂ he ↦ @t_sub ⟨(x₁, x₂), he⟩ (prod_sub ⟨h₁.1, h₂.2⟩)⟩ -theorem IsLocallyInjective_iff_openEmbedding {f : X → Y} : - IsLocallyInjective f ↔ OpenEmbedding (toPullbackDiag f) := by +theorem IsLocallyInjective_iff_isOpenEmbedding {f : X → Y} : + IsLocallyInjective f ↔ IsOpenEmbedding (toPullbackDiag f) := by rw [isLocallyInjective_iff_isOpen_diagonal, ← range_toPullbackDiag] exact ⟨fun h ↦ ⟨embedding_toPullbackDiag f, h⟩, fun h ↦ h.isOpen_range⟩ +@[deprecated (since := "2024-10-18")] +alias IsLocallyInjective_iff_openEmbedding := IsLocallyInjective_iff_isOpenEmbedding + theorem isLocallyInjective_iff_isOpenMap {f : X → Y} : IsLocallyInjective f ↔ IsOpenMap (toPullbackDiag f) := - IsLocallyInjective_iff_openEmbedding.trans - ⟨OpenEmbedding.isOpenMap, openEmbedding_of_continuous_injective_open + IsLocallyInjective_iff_isOpenEmbedding.trans + ⟨IsOpenEmbedding.isOpenMap, isOpenEmbedding_of_continuous_injective_open (embedding_toPullbackDiag f).continuous (injective_toPullbackDiag f)⟩ theorem discreteTopology_iff_locallyInjective (y : Y) : diff --git a/Mathlib/Topology/Separation.lean b/Mathlib/Topology/Separation.lean index 792e2e67b2544..1ba06927812eb 100644 --- a/Mathlib/Topology/Separation.lean +++ b/Mathlib/Topology/Separation.lean @@ -1544,7 +1544,7 @@ Hausdorff spaces: `f x ≠ f y`. We use this lemma to prove that topological spaces defined using `induced` are Hausdorff spaces. -* `separated_by_openEmbedding` says that for an open embedding `f : X → Y` of a Hausdorff space +* `separated_by_isOpenEmbedding` says that for an open embedding `f : X → Y` of a Hausdorff space `X`, the images of two distinct points `x y : X`, `x ≠ y` can be separated by open neighborhoods. We use this lemma to prove that topological spaces defined using `coinduced` are Hausdorff spaces. -/ @@ -1560,13 +1560,16 @@ theorem separated_by_continuous [TopologicalSpace Y] [T2Space Y] let ⟨u, v, uo, vo, xu, yv, uv⟩ := t2_separation h ⟨f ⁻¹' u, f ⁻¹' v, uo.preimage hf, vo.preimage hf, xu, yv, uv.preimage _⟩ -theorem separated_by_openEmbedding [TopologicalSpace Y] [T2Space X] - {f : X → Y} (hf : OpenEmbedding f) {x y : X} (h : x ≠ y) : +theorem separated_by_isOpenEmbedding [TopologicalSpace Y] [T2Space X] + {f : X → Y} (hf : IsOpenEmbedding f) {x y : X} (h : x ≠ y) : ∃ u v : Set Y, IsOpen u ∧ IsOpen v ∧ f x ∈ u ∧ f y ∈ v ∧ Disjoint u v := let ⟨u, v, uo, vo, xu, yv, uv⟩ := t2_separation h ⟨f '' u, f '' v, hf.isOpenMap _ uo, hf.isOpenMap _ vo, mem_image_of_mem _ xu, mem_image_of_mem _ yv, disjoint_image_of_injective hf.inj uv⟩ +@[deprecated (since := "2024-10-18")] +alias separated_by_openEmbedding := separated_by_isOpenEmbedding + instance {p : X → Prop} [T2Space X] : T2Space (Subtype p) := inferInstance instance Prod.t2Space [T2Space X] [TopologicalSpace Y] [T2Space Y] : T2Space (X × Y) := @@ -1591,10 +1594,10 @@ instance [T2Space X] [TopologicalSpace Y] [T2Space Y] : T2Space (X ⊕ Y) := by constructor rintro (x | x) (y | y) h - · exact separated_by_openEmbedding openEmbedding_inl <| ne_of_apply_ne _ h + · exact separated_by_isOpenEmbedding isOpenEmbedding_inl <| ne_of_apply_ne _ h · exact separated_by_continuous continuous_isLeft <| by simp · exact separated_by_continuous continuous_isLeft <| by simp - · exact separated_by_openEmbedding openEmbedding_inr <| ne_of_apply_ne _ h + · exact separated_by_isOpenEmbedding isOpenEmbedding_inr <| ne_of_apply_ne _ h instance Pi.t2Space {Y : X → Type v} [∀ a, TopologicalSpace (Y a)] [∀ a, T2Space (Y a)] : T2Space (∀ a, Y a) := @@ -1606,7 +1609,7 @@ instance Sigma.t2Space {ι} {X : ι → Type*} [∀ i, TopologicalSpace (X i)] [ rintro ⟨i, x⟩ ⟨j, y⟩ neq rcases eq_or_ne i j with (rfl | h) · replace neq : x ≠ y := ne_of_apply_ne _ neq - exact separated_by_openEmbedding openEmbedding_sigmaMk neq + exact separated_by_isOpenEmbedding isOpenEmbedding_sigmaMk neq · let _ := (⊥ : TopologicalSpace ι); have : DiscreteTopology ι := ⟨rfl⟩ exact separated_by_continuous (continuous_def.2 fun u _ => isOpen_sigma_fst_preimage u) h @@ -2597,7 +2600,7 @@ theorem loc_compact_Haus_tot_disc_of_zero_dim [TotallyDisconnectedSpace H] : let v : Set u := ((↑) : u → s) ⁻¹' V have : ((↑) : u → H) = ((↑) : s → H) ∘ ((↑) : u → s) := rfl have f0 : Embedding ((↑) : u → H) := embedding_subtype_val.comp embedding_subtype_val - have f1 : OpenEmbedding ((↑) : u → H) := by + have f1 : IsOpenEmbedding ((↑) : u → H) := by refine ⟨f0, ?_⟩ · have : Set.range ((↑) : u → H) = interior s := by rw [this, Set.range_comp, Subtype.range_coe, Subtype.image_preimage_coe] diff --git a/Mathlib/Topology/Sets/Opens.lean b/Mathlib/Topology/Sets/Opens.lean index bd142710c0990..1252c062af227 100644 --- a/Mathlib/Topology/Sets/Opens.lean +++ b/Mathlib/Topology/Sets/Opens.lean @@ -243,16 +243,22 @@ def frameMinimalAxioms : Frame.MinimalAxioms (Opens α) where instance instFrame : Frame (Opens α) := .ofMinimalAxioms frameMinimalAxioms -theorem openEmbedding' (U : Opens α) : OpenEmbedding (Subtype.val : U → α) := - U.isOpen.openEmbedding_subtype_val +theorem isOpenEmbedding' (U : Opens α) : IsOpenEmbedding (Subtype.val : U → α) := + U.isOpen.isOpenEmbedding_subtypeVal -theorem openEmbedding_of_le {U V : Opens α} (i : U ≤ V) : - OpenEmbedding (Set.inclusion <| SetLike.coe_subset_coe.2 i) := +@[deprecated (since := "2024-10-18")] +alias openEmbedding' := isOpenEmbedding' + +theorem isOpenEmbedding_of_le {U V : Opens α} (i : U ≤ V) : + IsOpenEmbedding (Set.inclusion <| SetLike.coe_subset_coe.2 i) := { toEmbedding := embedding_inclusion i isOpen_range := by rw [Set.range_inclusion i] exact U.isOpen.preimage continuous_subtype_val } +@[deprecated (since := "2024-10-18")] +alias openEmbedding_of_le := isOpenEmbedding_of_le + theorem not_nonempty_iff_eq_bot (U : Opens α) : ¬Set.Nonempty (U : Set α) ↔ U = ⊥ := by rw [← coe_inj, coe_bot, ← Set.not_nonempty_iff_eq_empty] diff --git a/Mathlib/Topology/Sheaves/LocalPredicate.lean b/Mathlib/Topology/Sheaves/LocalPredicate.lean index 8d32b5f87a342..bf6e5eb7f9b01 100644 --- a/Mathlib/Topology/Sheaves/LocalPredicate.lean +++ b/Mathlib/Topology/Sheaves/LocalPredicate.lean @@ -73,7 +73,7 @@ variable (X) @[simps!] def continuousPrelocal (T : TopCat.{v}) : PrelocalPredicate fun _ : X => T where pred {_} f := Continuous f - res {_ _} i _ h := Continuous.comp h (Opens.openEmbedding_of_le i.le).continuous + res {_ _} i _ h := Continuous.comp h (Opens.isOpenEmbedding_of_le i.le).continuous /-- Satisfying the inhabited linter. -/ instance inhabitedPrelocalPredicate (T : TopCat.{v}) : @@ -114,7 +114,7 @@ def continuousLocal (T : TopCat.{v}) : LocalPredicate fun _ : X => T := dsimp at w rw [continuous_iff_continuousAt] at w specialize w ⟨x, m⟩ - simpa using (Opens.openEmbedding_of_le i.le).continuousAt_iff.1 w } + simpa using (Opens.isOpenEmbedding_of_le i.le).continuousAt_iff.1 w } /-- Satisfying the inhabited linter. -/ instance inhabitedLocalPredicate (T : TopCat.{v}) : Inhabited (LocalPredicate fun _ : X => T) := diff --git a/Mathlib/Topology/Sheaves/SheafCondition/Sites.lean b/Mathlib/Topology/Sheaves/SheafCondition/Sites.lean index 07a4029efeb34..88541de424f6b 100644 --- a/Mathlib/Topology/Sheaves/SheafCondition/Sites.lean +++ b/Mathlib/Topology/Sheaves/SheafCondition/Sites.lean @@ -136,14 +136,14 @@ theorem coverDense_inducedFunctor {B : ι → Opens X} (h : Opens.IsBasis (Set.r end TopCat.Opens -section OpenEmbedding +section IsOpenEmbedding open TopCat.Presheaf Opposite variable {C : Type u} [Category.{v} C] variable {X Y : TopCat.{w}} {f : X ⟶ Y} {F : Y.Presheaf C} -theorem OpenEmbedding.compatiblePreserving (hf : OpenEmbedding f) : +theorem IsOpenEmbedding.compatiblePreserving (hf : IsOpenEmbedding f) : CompatiblePreserving (Opens.grothendieckTopology Y) hf.isOpenMap.functor := by haveI : Mono f := (TopCat.mono_iff_injective f).mpr hf.inj apply compatiblePreservingOfDownwardsClosed @@ -152,6 +152,9 @@ theorem OpenEmbedding.compatiblePreserving (hf : OpenEmbedding f) : obtain ⟨_, _, rfl⟩ := i.le h exact ⟨_, rfl⟩ +@[deprecated (since := "2024-10-18")] +alias OpenEmbedding.compatiblePreserving := IsOpenEmbedding.compatiblePreserving + theorem IsOpenMap.coverPreserving (hf : IsOpenMap f) : CoverPreserving (Opens.grothendieckTopology X) (Opens.grothendieckTopology Y) hf.functor := by constructor @@ -160,18 +163,24 @@ theorem IsOpenMap.coverPreserving (hf : IsOpenMap f) : exact ⟨_, hf.functor.map i, ⟨_, i, 𝟙 _, hV, rfl⟩, Set.mem_image_of_mem f hxV⟩ -lemma OpenEmbedding.functor_isContinuous (h : OpenEmbedding f) : +lemma IsOpenEmbedding.functor_isContinuous (h : IsOpenEmbedding f) : h.isOpenMap.functor.IsContinuous (Opens.grothendieckTopology X) (Opens.grothendieckTopology Y) := by apply Functor.isContinuous_of_coverPreserving · exact h.compatiblePreserving · exact h.isOpenMap.coverPreserving -theorem TopCat.Presheaf.isSheaf_of_openEmbedding (h : OpenEmbedding f) (hF : F.IsSheaf) : +@[deprecated (since := "2024-10-18")] +alias OpenEmbedding.functor_isContinuous := IsOpenEmbedding.functor_isContinuous + +theorem TopCat.Presheaf.isSheaf_of_isOpenEmbedding (h : IsOpenEmbedding f) (hF : F.IsSheaf) : IsSheaf (h.isOpenMap.functor.op ⋙ F) := by have := h.functor_isContinuous exact Functor.op_comp_isSheaf _ _ _ ⟨_, hF⟩ +@[deprecated (since := "2024-10-18")] +alias TopCat.Presheaf.isSheaf_of_openEmbedding := TopCat.Presheaf.isSheaf_of_isOpenEmbedding + variable (f) instance : RepresentablyFlat (Opens.map f) := by @@ -203,7 +212,7 @@ instance : (Opens.map f).IsContinuous (Opens.grothendieckTopology Y) · exact compatiblePreserving_opens_map f · exact coverPreserving_opens_map f -end OpenEmbedding +end IsOpenEmbedding namespace TopCat.Sheaf diff --git a/Mathlib/Topology/Sheaves/Stalks.lean b/Mathlib/Topology/Sheaves/Stalks.lean index 74df85eb39c03..125a7e70f0fc4 100644 --- a/Mathlib/Topology/Sheaves/Stalks.lean +++ b/Mathlib/Topology/Sheaves/Stalks.lean @@ -208,8 +208,8 @@ theorem comp (ℱ : X.Presheaf C) (f : X ⟶ Y) (g : Y ⟶ Z) (x : X) : ext simp [germ, stalkPushforward] -theorem stalkPushforward_iso_of_openEmbedding {f : X ⟶ Y} (hf : OpenEmbedding f) (F : X.Presheaf C) - (x : X) : IsIso (F.stalkPushforward _ f x) := by +theorem stalkPushforward_iso_of_isOpenEmbedding {f : X ⟶ Y} (hf : IsOpenEmbedding f) + (F : X.Presheaf C) (x : X) : IsIso (F.stalkPushforward _ f x) := by haveI := Functor.initial_of_adjunction (hf.isOpenMap.adjunctionNhds x) convert ((Functor.Final.colimitIso (hf.isOpenMap.functorNhds x).op @@ -234,6 +234,9 @@ theorem stalkPushforward_iso_of_openEmbedding {f : X ⟶ Y} (hf : OpenEmbedding refine ((homOfLE ?_).op : op (unop U) ⟶ _) exact Set.image_preimage_subset _ _ +@[deprecated (since := "2024-10-18")] +alias stalkPushforward_iso_of_openEmbedding := stalkPushforward_iso_of_isOpenEmbedding + end stalkPushforward section stalkPullback diff --git a/Mathlib/Topology/Sober.lean b/Mathlib/Topology/Sober.lean index 82c12d98222d6..06f9286a3e5d3 100644 --- a/Mathlib/Topology/Sober.lean +++ b/Mathlib/Topology/Sober.lean @@ -182,7 +182,7 @@ theorem ClosedEmbedding.quasiSober {f : α → β} (hf : ClosedEmbedding f) [Qua apply image_injective.mpr hf.inj rw [← hx.def, ← hf.closure_image_eq, image_singleton] -theorem OpenEmbedding.quasiSober {f : α → β} (hf : OpenEmbedding f) [QuasiSober β] : +theorem IsOpenEmbedding.quasiSober {f : α → β} (hf : IsOpenEmbedding f) [QuasiSober β] : QuasiSober α where sober hS hS' := by have hS'' := hS.image f hf.continuous.continuousOn @@ -205,6 +205,9 @@ theorem OpenEmbedding.quasiSober {f : α → β} (hf : OpenEmbedding f) [QuasiSo exact fun hy => ⟨fun h => hT.closure_eq ▸ closure_mono inter_subset_left h, fun h => subset_closure ⟨h, hy⟩⟩ +@[deprecated (since := "2024-10-18")] +alias OpenEmbedding.quasiSober := IsOpenEmbedding.quasiSober + /-- A space is quasi sober if it can be covered by open quasi sober subsets. -/ theorem quasiSober_of_open_cover (S : Set (Set α)) (hS : ∀ s : S, IsOpen (s : Set α)) [hS' : ∀ s : S, QuasiSober s] (hS'' : ⋃₀ S = ⊤) : QuasiSober α := by @@ -216,7 +219,7 @@ theorem quasiSober_of_open_cover (S : Set (Set α)) (hS : ∀ s : S, IsOpen (s : trivial haveI : QuasiSober U := hS' ⟨U, hU⟩ have H : IsPreirreducible ((↑) ⁻¹' t : Set U) := - h.2.preimage (hS ⟨U, hU⟩).openEmbedding_subtype_val + h.2.preimage (hS ⟨U, hU⟩).isOpenEmbedding_subtypeVal replace H : IsIrreducible ((↑) ⁻¹' t : Set U) := ⟨⟨⟨x, hU'⟩, by simpa using hx⟩, H⟩ use H.genericPoint have := continuous_subtype_val.closure_preimage_subset _ H.isGenericPoint_genericPoint_closure.mem diff --git a/Mathlib/Topology/Support.lean b/Mathlib/Topology/Support.lean index c0f19ee3e63a7..11a71da0311a5 100644 --- a/Mathlib/Topology/Support.lean +++ b/Mathlib/Topology/Support.lean @@ -257,7 +257,7 @@ theorem continuous_extend_one [TopologicalSpace β] {U : Set α'} (hU : IsOpen U continuous_of_mulTSupport fun x h ↦ by rw [show x = ↑(⟨x, Subtype.coe_image_subset _ _ (supp.mulTSupport_extend_one_subset continuous_subtype_val h)⟩ : U) by rfl, - ← (hU.openEmbedding_subtype_val).continuousAt_iff, extend_comp Subtype.val_injective] + ← (hU.isOpenEmbedding_subtypeVal).continuousAt_iff, extend_comp Subtype.val_injective] exact cont.continuousAt /-- If `f` has compact multiplicative support, then `f` tends to 1 at infinity. -/ diff --git a/scripts/no_lints_prime_decls.txt b/scripts/no_lints_prime_decls.txt index 671479377b629..e1817a2092239 100644 --- a/scripts/no_lints_prime_decls.txt +++ b/scripts/no_lints_prime_decls.txt @@ -3521,7 +3521,7 @@ ONote.NF.below_of_lt' ONote.nf_repr_split' ONote.NF.snd' ONote.split_eq_scale_split' -OpenEmbedding.tendsto_nhds_iff' +IsOpenEmbedding.tendsto_nhds_iff' openSegment_eq_image' openSegment_eq_Ioo' Option.bind_congr' @@ -4639,8 +4639,8 @@ toIocMod_sub_zsmul' toIocMod_zsmul_add' toIxxMod_total' TopCat.GlueData.preimage_image_eq_image' -TopCat.openEmbedding_iff_comp_isIso' -TopCat.openEmbedding_iff_isIso_comp' +TopCat.isOpenEmbedding_iff_comp_isIso' +TopCat.isOpenEmbedding_iff_isIso_comp' TopCat.Presheaf.germ_stalkSpecializes' TopCat.Presheaf.pushforward_eq' TopCat.Presheaf.pushforward_map_app' @@ -4650,7 +4650,7 @@ TopologicalSpace.Opens.coe_inclusion' TopologicalSpace.Opens.map_comp_obj' TopologicalSpace.Opens.map_functor_eq' TopologicalSpace.Opens.map_id_obj' -TopologicalSpace.Opens.openEmbedding' +TopologicalSpace.Opens.isOpenEmbedding' TopologicalSpace.Opens.set_range_forget_map_inclusion' TopologicalSpace.Opens.set_range_inclusion' TopologicalSpace.SecondCountableTopology.mk' From b9efa9d8b132e603ec517380352c5b5a29a25e88 Mon Sep 17 00:00:00 2001 From: peabrainiac Date: Fri, 18 Oct 2024 21:51:41 +0000 Subject: [PATCH 166/505] feat(Topology/Connected/PathConnected): some lemmas for `pathComponentIn` (#16983) Adds some basic lemmas about `pathComponentIn`, analogous to existing lemmas about `pathComponent` and `connectedComponentIn`. Co-authored-by: peabrainiac <43812953+peabrainiac@users.noreply.github.com> --- Mathlib/Topology/Connected/PathConnected.lean | 33 +++++++++++++++++++ 1 file changed, 33 insertions(+) diff --git a/Mathlib/Topology/Connected/PathConnected.lean b/Mathlib/Topology/Connected/PathConnected.lean index 3a8377f938164..b9a49ba4fcd65 100644 --- a/Mathlib/Topology/Connected/PathConnected.lean +++ b/Mathlib/Topology/Connected/PathConnected.lean @@ -848,6 +848,24 @@ theorem Joined.mem_pathComponent (hyz : Joined y z) (hxy : y ∈ pathComponent x z ∈ pathComponent x := hxy.trans hyz +theorem mem_pathComponentIn_self (h : x ∈ F) : x ∈ pathComponentIn x F := + JoinedIn.refl h + +theorem pathComponentIn_subset : pathComponentIn x F ⊆ F := + fun _ hy ↦ hy.target_mem + +theorem pathComponentIn_nonempty_iff : (pathComponentIn x F).Nonempty ↔ x ∈ F := + ⟨fun ⟨_, ⟨γ, hγ⟩⟩ ↦ γ.source ▸ hγ 0, fun hx ↦ ⟨x, mem_pathComponentIn_self hx⟩⟩ + +theorem pathComponentIn_congr (h : x ∈ pathComponentIn y F) : + pathComponentIn x F = pathComponentIn y F := by + ext; exact ⟨h.trans, h.symm.trans⟩ + +@[gcongr] +theorem pathComponentIn_mono {G : Set X} (h : F ⊆ G) : + pathComponentIn x F ⊆ pathComponentIn x G := + fun _ ⟨γ, hγ⟩ ↦ ⟨γ, fun t ↦ h (hγ t)⟩ + /-! ### Path connected sets -/ @@ -913,11 +931,26 @@ theorem IsPathConnected.mem_pathComponent (h : IsPathConnected F) (x_in : x ∈ theorem IsPathConnected.subset_pathComponent (h : IsPathConnected F) (x_in : x ∈ F) : F ⊆ pathComponent x := fun _y y_in => h.mem_pathComponent x_in y_in +theorem IsPathConnected.subset_pathComponentIn {s : Set X} (hs : IsPathConnected s) + (hxs : x ∈ s) (hsF : s ⊆ F) : s ⊆ pathComponentIn x F := + fun y hys ↦ (hs.joinedIn x hxs y hys).mono hsF + theorem isPathConnected_singleton (x : X) : IsPathConnected ({x} : Set X) := by refine ⟨x, rfl, ?_⟩ rintro y rfl exact JoinedIn.refl rfl +theorem isPathConnected_pathComponentIn (h : x ∈ F) : IsPathConnected (pathComponentIn x F) := + ⟨x, mem_pathComponentIn_self h, fun ⟨γ, hγ⟩ ↦ by + refine ⟨γ, fun t ↦ + ⟨(γ.truncateOfLE t.2.1).cast (γ.extend_zero.symm) (γ.extend_extends' t).symm, fun t' ↦ ?_⟩⟩ + dsimp [Path.truncateOfLE, Path.truncate] + exact γ.extend_extends' ⟨min (max t'.1 0) t.1, by simp [t.2.1, t.2.2]⟩ ▸ hγ _⟩ + +theorem isPathConnected_pathComponent : IsPathConnected (pathComponent x) := by + rw [← pathComponentIn_univ] + exact isPathConnected_pathComponentIn (mem_univ x) + theorem IsPathConnected.union {U V : Set X} (hU : IsPathConnected U) (hV : IsPathConnected V) (hUV : (U ∩ V).Nonempty) : IsPathConnected (U ∪ V) := by rcases hUV with ⟨x, xU, xV⟩ From fa8e2105aa86b708d0b03e08ba31c88c1a1f982c Mon Sep 17 00:00:00 2001 From: mathlib4-update-dependencies-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Fri, 18 Oct 2024 21:51:43 +0000 Subject: [PATCH 167/505] chore: update Mathlib dependencies 2024-10-18 (#17929) This PR updates the Mathlib dependencies. --- lake-manifest.json | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lake-manifest.json b/lake-manifest.json index c93ba44db832e..f8536ac19fc8c 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,7 +5,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "c521f0185f4dd42b6aa4898010d5ba5357c57c9f", + "rev": "9e3d0d810e9021767c360756f066574f72e8fcce", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", From 992ec73d5d2540c7d7bd628b704ef49909a13835 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sat, 19 Oct 2024 00:06:23 +0000 Subject: [PATCH 168/505] feat: generalize `LinearMap.exists_rightInverse_of_surjective` to projective modules (#17560) Previously it only held in vector spaces. --- Mathlib/Algebra/Module/Projective.lean | 6 +++++- Mathlib/LinearAlgebra/Basis/VectorSpace.lean | 9 --------- Mathlib/LinearAlgebra/Dimension/LinearMap.lean | 1 + Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean | 1 + Mathlib/Topology/Algebra/SeparationQuotient.lean | 1 + 5 files changed, 8 insertions(+), 10 deletions(-) diff --git a/Mathlib/Algebra/Module/Projective.lean b/Mathlib/Algebra/Module/Projective.lean index b4387f01f63c9..7a68117b7f456 100644 --- a/Mathlib/Algebra/Module/Projective.lean +++ b/Mathlib/Algebra/Module/Projective.lean @@ -91,7 +91,7 @@ theorem projective_def' : /-- A projective R-module has the property that maps from it lift along surjections. -/ theorem projective_lifting_property [h : Projective R P] (f : M →ₗ[R] N) (g : P →ₗ[R] N) - (hf : Function.Surjective f) : ∃ h : P →ₗ[R] M, f.comp h = g := by + (hf : Function.Surjective f) : ∃ h : P →ₗ[R] M, f ∘ₗ h = g := by /- Here's the first step of the proof. Recall that `X →₀ R` is Lean's way of talking about the free `R`-module @@ -110,6 +110,10 @@ theorem projective_lifting_property [h : Projective R P] (f : M →ₗ[R] N) (g conv_rhs => rw [← hs p] simp [φ, Finsupp.linearCombination_apply, Function.surjInv_eq hf, map_finsupp_sum] +theorem _root_.LinearMap.exists_rightInverse_of_surjective [Projective R P] + (f : M →ₗ[R] P) (hf_surj : range f = ⊤) : ∃ g : P →ₗ[R] M, f ∘ₗ g = LinearMap.id := + projective_lifting_property f (.id : P →ₗ[R] P) (LinearMap.range_eq_top.1 hf_surj) + /-- A module which satisfies the universal property is projective: If all surjections of `R`-modules `(P →₀ R) →ₗ[R] P` have `R`-linear left inverse maps, then `P` is projective. -/ diff --git a/Mathlib/LinearAlgebra/Basis/VectorSpace.lean b/Mathlib/LinearAlgebra/Basis/VectorSpace.lean index e7a219e569dab..c16a5b57e71ea 100644 --- a/Mathlib/LinearAlgebra/Basis/VectorSpace.lean +++ b/Mathlib/LinearAlgebra/Basis/VectorSpace.lean @@ -266,15 +266,6 @@ theorem Submodule.exists_isCompl (p : Submodule K V) : ∃ q : Submodule K V, Is instance Submodule.complementedLattice : ComplementedLattice (Submodule K V) := ⟨Submodule.exists_isCompl⟩ -theorem LinearMap.exists_rightInverse_of_surjective (f : V →ₗ[K] V') (hf_surj : range f = ⊤) : - ∃ g : V' →ₗ[K] V, f.comp g = LinearMap.id := by - let C := Basis.ofVectorSpaceIndex K V' - let hC := Basis.ofVectorSpace K V' - haveI : Inhabited V := ⟨0⟩ - refine ⟨(hC.constr ℕ : _ → _) (C.restrict (invFun f)), hC.ext fun c => ?_⟩ - rw [LinearMap.comp_apply, hC.constr_basis] - simp [hC, rightInverse_invFun (LinearMap.range_eq_top.1 hf_surj) c] - /-- Any linear map `f : p →ₗ[K] V'` defined on a subspace `p` can be extended to the whole space. -/ theorem LinearMap.exists_extend {p : Submodule K V} (f : p →ₗ[K] V') : diff --git a/Mathlib/LinearAlgebra/Dimension/LinearMap.lean b/Mathlib/LinearAlgebra/Dimension/LinearMap.lean index ca2cb869a8ec5..d9f144980e339 100644 --- a/Mathlib/LinearAlgebra/Dimension/LinearMap.lean +++ b/Mathlib/LinearAlgebra/Dimension/LinearMap.lean @@ -3,6 +3,7 @@ Copyright (c) 2019 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl -/ +import Mathlib.Algebra.Module.Projective import Mathlib.LinearAlgebra.Dimension.DivisionRing import Mathlib.LinearAlgebra.Dimension.FreeAndStrongRankCondition diff --git a/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean b/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean index 172f5f1c0d87a..372c3c0e74552 100644 --- a/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean +++ b/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean @@ -3,6 +3,7 @@ Copyright (c) 2019 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes -/ +import Mathlib.Algebra.Module.Projective import Mathlib.FieldTheory.Finiteness /-! diff --git a/Mathlib/Topology/Algebra/SeparationQuotient.lean b/Mathlib/Topology/Algebra/SeparationQuotient.lean index 83775bba6cbe3..8a99f4dc7ef34 100644 --- a/Mathlib/Topology/Algebra/SeparationQuotient.lean +++ b/Mathlib/Topology/Algebra/SeparationQuotient.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import Mathlib.LinearAlgebra.Basis.VectorSpace +import Mathlib.Algebra.Module.Projective import Mathlib.Topology.Algebra.Module.Basic import Mathlib.Topology.Maps.OpenQuotient From 6e38c6649191d8bf77deb739de42ae8df4581c0d Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sat, 19 Oct 2024 00:45:04 +0000 Subject: [PATCH 169/505] feat: algebraic properties of `LinearMap.compMultilinear` (#17932) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This also demotes `LinearMap.compAlternatingMap` to a plain function, for consistency. To make up for this loss, it introduces `LinearMap.compAlternatingMapₗ` as the stronger linear version. --- Mathlib/LinearAlgebra/Alternating/Basic.lean | 56 ++++++++++++++++---- Mathlib/LinearAlgebra/Multilinear/Basic.lean | 47 ++++++++++++++-- 2 files changed, 88 insertions(+), 15 deletions(-) diff --git a/Mathlib/LinearAlgebra/Alternating/Basic.lean b/Mathlib/LinearAlgebra/Alternating/Basic.lean index 994b3b3d983db..f3a7aafc8973a 100644 --- a/Mathlib/LinearAlgebra/Alternating/Basic.lean +++ b/Mathlib/LinearAlgebra/Alternating/Basic.lean @@ -407,19 +407,12 @@ end AlternatingMap namespace LinearMap -variable {N₂ : Type*} [AddCommMonoid N₂] [Module R N₂] +variable {S : Type*} {N₂ : Type*} [AddCommMonoid N₂] [Module R N₂] /-- Composing an alternating map with a linear map on the left gives again an alternating map. -/ -def compAlternatingMap (g : N →ₗ[R] N₂) : (M [⋀^ι]→ₗ[R] N) →+ (M [⋀^ι]→ₗ[R] N₂) where - toFun f := - { g.compMultilinearMap (f : MultilinearMap R (fun _ : ι => M) N) with - map_eq_zero_of_eq' := fun v i j h hij => by simp [f.map_eq_zero_of_eq v h hij] } - map_zero' := by - ext - simp - map_add' a b := by - ext - simp +def compAlternatingMap (g : N →ₗ[R] N₂) (f : M [⋀^ι]→ₗ[R] N) : M [⋀^ι]→ₗ[R] N₂ where + __ := g.compMultilinearMap (f : MultilinearMap R (fun _ : ι => M) N) + map_eq_zero_of_eq' v i j h hij := by simp [f.map_eq_zero_of_eq v h hij] @[simp] theorem coe_compAlternatingMap (g : N →ₗ[R] N₂) (f : M [⋀^ι]→ₗ[R] N) : @@ -431,6 +424,47 @@ theorem compAlternatingMap_apply (g : N →ₗ[R] N₂) (f : M [⋀^ι]→ₗ[R] g.compAlternatingMap f m = g (f m) := rfl +@[simp] +theorem compAlternatingMap_zero (g : N →ₗ[R] N₂) : + g.compAlternatingMap (0 : M [⋀^ι]→ₗ[R] N) = 0 := + AlternatingMap.ext fun _ => map_zero g + +@[simp] +theorem zero_compAlternatingMap (f: M [⋀^ι]→ₗ[R] N) : + (0 : N →ₗ[R] N₂).compAlternatingMap f = 0 := rfl + +@[simp] +theorem compAlternatingMap_add (g : N →ₗ[R] N₂) (f₁ f₂ : M [⋀^ι]→ₗ[R] N) : + g.compAlternatingMap (f₁ + f₂) = g.compAlternatingMap f₁ + g.compAlternatingMap f₂ := + AlternatingMap.ext fun _ => map_add g _ _ + +@[simp] +theorem add_compAlternatingMap (g₁ g₂ : N →ₗ[R] N₂) (f: M [⋀^ι]→ₗ[R] N) : + (g₁ + g₂).compAlternatingMap f = g₁.compAlternatingMap f + g₂.compAlternatingMap f := rfl + +@[simp] +theorem compAlternatingMap_smul [Monoid S] [DistribMulAction S N] [DistribMulAction S N₂] + [SMulCommClass R S N] [SMulCommClass R S N₂] [CompatibleSMul N N₂ S R] + (g : N →ₗ[R] N₂) (s : S) (f : M [⋀^ι]→ₗ[R] N) : + g.compAlternatingMap (s • f) = s • g.compAlternatingMap f := + AlternatingMap.ext fun _ => g.map_smul_of_tower _ _ + +@[simp] +theorem smul_compAlternatingMap [Monoid S] [DistribMulAction S N₂] [SMulCommClass R S N₂] + (g : N →ₗ[R] N₂) (s : S) (f : M [⋀^ι]→ₗ[R] N) : + (s • g).compAlternatingMap f = s • g.compAlternatingMap f := rfl + +variable (S) in +/-- `LinearMap.compAlternatingMap` as an `S`-linear map. -/ +@[simps] +def compAlternatingMapₗ [Semiring S] [Module S N] [Module S N₂] + [SMulCommClass R S N] [SMulCommClass R S N₂] [LinearMap.CompatibleSMul N N₂ S R] + (g : N →ₗ[R] N₂) : + (M [⋀^ι]→ₗ[R] N) →ₗ[S] (M [⋀^ι]→ₗ[R] N₂) where + toFun := g.compAlternatingMap + map_add' := g.compAlternatingMap_add + map_smul' := g.compAlternatingMap_smul + theorem smulRight_eq_comp {R M₁ M₂ ι : Type*} [CommSemiring R] [AddCommMonoid M₁] [AddCommMonoid M₂] [Module R M₁] [Module R M₂] (f : M₁ [⋀^ι]→ₗ[R] R) (z : M₂) : f.smulRight z = (LinearMap.id.smulRight z).compAlternatingMap f := diff --git a/Mathlib/LinearAlgebra/Multilinear/Basic.lean b/Mathlib/LinearAlgebra/Multilinear/Basic.lean index 5fa2ebd1c7ff4..8db7f46936b45 100644 --- a/Mathlib/LinearAlgebra/Multilinear/Basic.lean +++ b/Mathlib/LinearAlgebra/Multilinear/Basic.lean @@ -787,6 +787,36 @@ theorem compMultilinearMap_apply (g : M₂ →ₗ[R] M₃) (f : MultilinearMap R g.compMultilinearMap f m = g (f m) := rfl +@[simp] +theorem compMultilinearMap_zero (g : M₂ →ₗ[R] M₃) : + g.compMultilinearMap (0 : MultilinearMap R M₁ M₂) = 0 := + MultilinearMap.ext fun _ => map_zero g + +@[simp] +theorem zero_compMultilinearMap (f: MultilinearMap R M₁ M₂) : + (0 : M₂ →ₗ[R] M₃).compMultilinearMap f = 0 := rfl + +@[simp] +theorem compMultilinearMap_add (g : M₂ →ₗ[R] M₃) (f₁ f₂ : MultilinearMap R M₁ M₂) : + g.compMultilinearMap (f₁ + f₂) = g.compMultilinearMap f₁ + g.compMultilinearMap f₂ := + MultilinearMap.ext fun _ => map_add g _ _ + +@[simp] +theorem add_compMultilinearMap (g₁ g₂ : M₂ →ₗ[R] M₃) (f: MultilinearMap R M₁ M₂) : + (g₁ + g₂).compMultilinearMap f = g₁.compMultilinearMap f + g₂.compMultilinearMap f := rfl + +@[simp] +theorem compMultilinearMap_smul [Monoid S] [DistribMulAction S M₂] [DistribMulAction S M₃] + [SMulCommClass R S M₂] [SMulCommClass R S M₃] [CompatibleSMul M₂ M₃ S R] + (g : M₂ →ₗ[R] M₃) (s : S) (f : MultilinearMap R M₁ M₂) : + g.compMultilinearMap (s • f) = s • g.compMultilinearMap f := + MultilinearMap.ext fun _ => g.map_smul_of_tower _ _ + +@[simp] +theorem smul_compMultilinearMap [Monoid S] [DistribMulAction S M₃] [SMulCommClass R S M₃] + (g : M₂ →ₗ[R] M₃) (s : S) (f : MultilinearMap R M₁ M₂) : + (s • g).compMultilinearMap f = s • g.compMultilinearMap f := rfl + /-- The multilinear version of `LinearMap.subtype_comp_codRestrict` -/ @[simp] theorem subtype_compMultilinearMap_codRestrict (f : MultilinearMap R M₁ M₂) (p : Submodule R M₂) @@ -835,12 +865,23 @@ instance : Module S (MultilinearMap R M₁ M₂) := instance [NoZeroSMulDivisors S M₂] : NoZeroSMulDivisors S (MultilinearMap R M₁ M₂) := coe_injective.noZeroSMulDivisors _ rfl coe_smul +variable [AddCommMonoid M₃] [Module S M₃] [Module R M₃] [SMulCommClass R S M₃] + +variable (S) in +/-- `LinearMap.compMultilinearMap` as an `S`-linear map. -/ +@[simps] +def _root_.LinearMap.compMultilinearMapₗ [Semiring S] [Module S M₂] [Module S M₃] + [SMulCommClass R S M₂] [SMulCommClass R S M₃] [LinearMap.CompatibleSMul M₂ M₃ S R] + (g : M₂ →ₗ[R] M₃) : + MultilinearMap R M₁ M₂ →ₗ[S] MultilinearMap R M₁ M₃ where + toFun := g.compMultilinearMap + map_add' := g.compMultilinearMap_add + map_smul' := g.compMultilinearMap_smul + variable (R S M₁ M₂ M₃) section OfSubsingleton -variable [AddCommMonoid M₃] [Module S M₃] [Module R M₃] [SMulCommClass R S M₃] - /-- Linear equivalence between linear maps `M₂ →ₗ[R] M₃` and one-multilinear maps `MultilinearMap R (fun _ : ι ↦ M₂) M₃`. -/ @[simps (config := { simpRhs := true })] @@ -904,8 +945,6 @@ def constLinearEquivOfIsEmpty [IsEmpty ι] : M₂ ≃ₗ[S] MultilinearMap R M left_inv _ := rfl right_inv f := ext fun _ => MultilinearMap.congr_arg f <| Subsingleton.elim _ _ -variable [AddCommMonoid M₃] [Module R M₃] [Module S M₃] [SMulCommClass R S M₃] - /-- `MultilinearMap.domDomCongr` as a `LinearEquiv`. -/ @[simps apply symm_apply] def domDomCongrLinearEquiv {ι₁ ι₂} (σ : ι₁ ≃ ι₂) : From f6b96180cc040c8fdf7eb10746000229ef5eba25 Mon Sep 17 00:00:00 2001 From: mathlib4-update-dependencies-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Sat, 19 Oct 2024 01:19:58 +0000 Subject: [PATCH 170/505] chore: update Mathlib dependencies 2024-10-19 (#17933) This PR updates the Mathlib dependencies. --- lake-manifest.json | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lake-manifest.json b/lake-manifest.json index f8536ac19fc8c..6f116ba1524dd 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,7 +5,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "9e3d0d810e9021767c360756f066574f72e8fcce", + "rev": "1e0bf50b357069e1d658512a579a5faac6587c40", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", From bd5ba6b8fd201d0165ca6068e0a9437a56c4b58f Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Sat, 19 Oct 2024 02:20:12 +0000 Subject: [PATCH 171/505] chore(SetTheory/Ordinal/Basic): deprecate duplicate `Nat.cast` lemmas (#17857) --- Mathlib/SetTheory/Ordinal/Arithmetic.lean | 37 +++++++++++------------ Mathlib/SetTheory/Ordinal/Notation.lean | 18 +++++------ 2 files changed, 27 insertions(+), 28 deletions(-) diff --git a/Mathlib/SetTheory/Ordinal/Arithmetic.lean b/Mathlib/SetTheory/Ordinal/Arithmetic.lean index 90d58666720d4..b43f781637c4a 100644 --- a/Mathlib/SetTheory/Ordinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Ordinal/Arithmetic.lean @@ -2099,6 +2099,10 @@ theorem Ordinal.not_bddAbove_compl_of_small (s : Set Ordinal.{u}) [hs : Small.{u namespace Ordinal +instance instCharZero : CharZero Ordinal := by + refine ⟨fun a b h ↦ ?_⟩ + rwa [← Cardinal.ord_nat, ← Cardinal.ord_nat, Cardinal.ord_inj, Nat.cast_inj] at h + @[simp] theorem one_add_natCast (m : ℕ) : 1 + (m : Ordinal) = succ m := by rw [← Nat.cast_one, ← Nat.cast_add, add_comm] @@ -2121,42 +2125,37 @@ theorem natCast_mul (m : ℕ) : ∀ n : ℕ, ((m * n : ℕ) : Ordinal) = m * n @[deprecated (since := "2024-04-17")] alias nat_cast_mul := natCast_mul -/-- Alias of `Nat.cast_le`, specialized to `Ordinal` --/ -theorem natCast_le {m n : ℕ} : (m : Ordinal) ≤ n ↔ m ≤ n := by - rw [← Cardinal.ord_nat, ← Cardinal.ord_nat, Cardinal.ord_le_ord, Cardinal.natCast_le] +@[deprecated Nat.cast_le (since := "2024-10-17")] +theorem natCast_le {m n : ℕ} : (m : Ordinal) ≤ n ↔ m ≤ n := Nat.cast_le @[deprecated (since := "2024-04-17")] alias nat_cast_le := natCast_le -/-- Alias of `Nat.cast_inj`, specialized to `Ordinal` --/ -theorem natCast_inj {m n : ℕ} : (m : Ordinal) = n ↔ m = n := by - simp only [le_antisymm_iff, natCast_le] +@[deprecated Nat.cast_inj (since := "2024-10-17")] +theorem natCast_inj {m n : ℕ} : (m : Ordinal) = n ↔ m = n := Nat.cast_inj @[deprecated (since := "2024-04-17")] alias nat_cast_inj := natCast_inj -instance charZero : CharZero Ordinal where - cast_injective _ _ := natCast_inj.mp - -/-- Alias of `Nat.cast_lt`, specialized to `Ordinal` --/ +@[deprecated Nat.cast_lt (since := "2024-10-17")] theorem natCast_lt {m n : ℕ} : (m : Ordinal) < n ↔ m < n := Nat.cast_lt @[deprecated (since := "2024-04-17")] alias nat_cast_lt := natCast_lt -/-- Alias of `Nat.cast_eq_zero`, specialized to `Ordinal` --/ +@[deprecated Nat.cast_eq_zero (since := "2024-10-17")] theorem natCast_eq_zero {n : ℕ} : (n : Ordinal) = 0 ↔ n = 0 := Nat.cast_eq_zero @[deprecated (since := "2024-04-17")] alias nat_cast_eq_zero := natCast_eq_zero -/-- Alias of `Nat.cast_eq_zero`, specialized to `Ordinal` --/ +@[deprecated Nat.cast_ne_zero (since := "2024-10-17")] theorem natCast_ne_zero {n : ℕ} : (n : Ordinal) ≠ 0 ↔ n ≠ 0 := Nat.cast_ne_zero @[deprecated (since := "2024-04-17")] alias nat_cast_ne_zero := natCast_ne_zero -/-- Alias of `Nat.cast_pos'`, specialized to `Ordinal` --/ +@[deprecated Nat.cast_pos' (since := "2024-10-17")] theorem natCast_pos {n : ℕ} : (0 : Ordinal) < n ↔ 0 < n := Nat.cast_pos' @[deprecated (since := "2024-04-17")] @@ -2165,10 +2164,10 @@ alias nat_cast_pos := natCast_pos @[simp, norm_cast] theorem natCast_sub (m n : ℕ) : ((m - n : ℕ) : Ordinal) = m - n := by rcases le_total m n with h | h - · rw [tsub_eq_zero_iff_le.2 h, Ordinal.sub_eq_zero_iff_le.2 (natCast_le.2 h)] + · rw [tsub_eq_zero_iff_le.2 h, Ordinal.sub_eq_zero_iff_le.2 (Nat.cast_le.2 h)] rfl · apply (add_left_cancel n).1 - rw [← Nat.cast_add, add_tsub_cancel_of_le h, Ordinal.add_sub_cancel_of_le (natCast_le.2 h)] + rw [← Nat.cast_add, add_tsub_cancel_of_le h, Ordinal.add_sub_cancel_of_le (Nat.cast_le.2 h)] @[deprecated (since := "2024-04-17")] alias nat_cast_sub := natCast_sub @@ -2177,12 +2176,12 @@ alias nat_cast_sub := natCast_sub theorem natCast_div (m n : ℕ) : ((m / n : ℕ) : Ordinal) = m / n := by rcases eq_or_ne n 0 with (rfl | hn) · simp - · have hn' := natCast_ne_zero.2 hn + · have hn' : (n : Ordinal) ≠ 0 := Nat.cast_ne_zero.2 hn apply le_antisymm - · rw [le_div hn', ← natCast_mul, natCast_le, mul_comm] + · rw [le_div hn', ← natCast_mul, Nat.cast_le, mul_comm] apply Nat.div_mul_le_self - · rw [div_le hn', ← add_one_eq_succ, ← Nat.cast_succ, ← natCast_mul, natCast_lt, mul_comm, ← - Nat.div_lt_iff_lt_mul (Nat.pos_of_ne_zero hn)] + · rw [div_le hn', ← add_one_eq_succ, ← Nat.cast_succ, ← natCast_mul, Nat.cast_lt, mul_comm, + ← Nat.div_lt_iff_lt_mul (Nat.pos_of_ne_zero hn)] apply Nat.lt_succ_self @[deprecated (since := "2024-04-17")] diff --git a/Mathlib/SetTheory/Ordinal/Notation.lean b/Mathlib/SetTheory/Ordinal/Notation.lean index 050c6b4f8f653..bab9ce4002c72 100644 --- a/Mathlib/SetTheory/Ordinal/Notation.lean +++ b/Mathlib/SetTheory/Ordinal/Notation.lean @@ -139,7 +139,7 @@ theorem repr_one : repr (ofNat 1) = (1 : ℕ) := repr_ofNat 1 theorem omega0_le_oadd (e n a) : ω ^ repr e ≤ repr (oadd e n a) := by refine le_trans ?_ (le_add_right _ _) - simpa using (Ordinal.mul_le_mul_iff_left <| opow_pos (repr e) omega0_pos).2 (natCast_le.2 n.2) + simpa using (Ordinal.mul_le_mul_iff_left <| opow_pos (repr e) omega0_pos).2 (Nat.cast_le.2 n.2) @[deprecated (since := "2024-09-30")] alias omega_le_oadd := omega0_le_oadd @@ -275,7 +275,7 @@ theorem oadd_lt_oadd_2 {e o₁ o₂ : ONote} {n₁ n₂ : ℕ+} (h₁ : NF (oadd oadd e n₁ o₁ < oadd e n₂ o₂ := by simp only [lt_def, repr] refine lt_of_lt_of_le ((add_lt_add_iff_left _).2 h₁.snd'.repr_lt) (le_trans ?_ (le_add_right _ _)) - rwa [← mul_succ,Ordinal.mul_le_mul_iff_left (opow_pos _ omega0_pos), succ_le_iff, natCast_lt] + rwa [← mul_succ,Ordinal.mul_le_mul_iff_left (opow_pos _ omega0_pos), succ_le_iff, Nat.cast_lt] theorem oadd_lt_oadd_3 {e n a₁ a₂} (h : a₁ < a₂) : oadd e n a₁ < oadd e n a₂ := by rw [lt_def]; unfold repr @@ -451,7 +451,7 @@ theorem repr_add : ∀ (o₁ o₂) [NF o₁] [NF o₂], repr (o₁ + o₂) = rep cases he' : e' <;> simp only [he', zero_def, opow_zero, repr, gt_iff_lt] at this ⊢ <;> exact lt_of_le_of_lt (le_add_right _ _) this · simpa using (Ordinal.mul_le_mul_iff_left <| opow_pos (repr e') omega0_pos).2 - (natCast_le.2 n'.pos) + (Nat.cast_le.2 n'.pos) · rw [ee, ← add_assoc, ← mul_add] theorem sub_nfBelow : ∀ {o₁ o₂ b}, NFBelow o₁ b → NF o₂ → NFBelow (o₁ - o₂) b @@ -507,7 +507,7 @@ theorem repr_sub : ∀ (o₁ o₂) [NF o₁] [NF o₂], repr (o₁ - o₂) = rep refine (Ordinal.sub_eq_of_add_eq <| add_absorp h₂.snd'.repr_lt <| le_trans ?_ (le_add_right _ _)).symm - simpa using mul_le_mul_left' (natCast_le.2 <| Nat.succ_pos _) _ + exact Ordinal.le_mul_left _ (Nat.cast_lt.2 <| Nat.succ_pos _) · exact (Ordinal.sub_eq_of_add_eq <| add_absorp (h₂.below_of_lt ee).repr_lt <| omega0_le_oadd _ _ _).symm @@ -564,7 +564,7 @@ theorem repr_mul : ∀ (o₁ o₂) [NF o₁] [NF o₂], repr (o₁ * o₂) = rep simp [(· * ·)] have ao : repr a₁ + ω ^ repr e₁ * (n₁ : ℕ) = ω ^ repr e₁ * (n₁ : ℕ) := by apply add_absorp h₁.snd'.repr_lt - simpa using (Ordinal.mul_le_mul_iff_left <| opow_pos _ omega0_pos).2 (natCast_le.2 n₁.2) + simpa using (Ordinal.mul_le_mul_iff_left <| opow_pos _ omega0_pos).2 (Nat.cast_le.2 n₁.2) by_cases e0 : e₂ = 0 · cases' Nat.exists_eq_succ_of_ne_zero n₂.ne_zero with x xe simp only [e0, repr, PNat.mul_coe, natCast_mul, opow_zero, one_mul] @@ -578,7 +578,7 @@ theorem repr_mul : ∀ (o₁ o₂) [NF o₁] [NF o₂], repr (o₁ * o₂) = rep congr 2 have := mt repr_inj.1 e0 rw [add_mul_limit ao (isLimit_opow_left isLimit_omega0 this), mul_assoc, - mul_omega0_dvd (natCast_pos.2 n₁.pos) (nat_lt_omega0 _)] + mul_omega0_dvd (Nat.cast_pos'.2 n₁.pos) (nat_lt_omega0 _)] simpa using opow_dvd_opow ω (one_le_iff_ne_zero.2 this) /-- Calculate division and remainder of `o` mod `ω`: @@ -857,7 +857,7 @@ theorem repr_opow_aux₂ {a0 a'} [N0 : NF a0] [Na' : NF a'] (m : ℕ) (d : ω · exact lt_of_lt_of_le Rl (opow_le_opow_right omega0_pos <| - mul_le_mul_left' (succ_le_succ_iff.2 (natCast_le.2 (le_of_lt k.lt_succ_self))) _) + mul_le_mul_left' (succ_le_succ_iff.2 (Nat.cast_le.2 (le_of_lt k.lt_succ_self))) _) calc (ω0 ^ (k.succ : Ordinal)) * α' + R' _ = (ω0 ^ succ (k : Ordinal)) * α' + ((ω0 ^ (k : Ordinal)) * α' * m + R) := by @@ -869,12 +869,12 @@ theorem repr_opow_aux₂ {a0 a'} [N0 : NF a0] [Na' : NF a'] (m : ℕ) (d : ω dvd_add (dvd_mul_of_dvd_left (by simpa using opow_dvd_opow ω (one_le_iff_ne_zero.2 e0)) _) d rw [mul_add (ω0 ^ (k : Ordinal)), add_assoc, ← mul_assoc, ← opow_succ, add_mul_limit _ (isLimit_iff_omega0_dvd.2 ⟨ne_of_gt α0, αd⟩), mul_assoc, - @mul_omega0_dvd n (natCast_pos.2 n.pos) (nat_lt_omega0 _) _ αd] + @mul_omega0_dvd n (Nat.cast_pos'.2 n.pos) (nat_lt_omega0 _) _ αd] apply @add_absorp _ (repr a0 * succ ↑k) · refine principal_add_omega0_opow _ ?_ Rl rw [opow_mul, opow_succ, Ordinal.mul_lt_mul_iff_left ω00] exact No.snd'.repr_lt - · have := mul_le_mul_left' (one_le_iff_pos.2 <| natCast_pos.2 n.pos) (ω0 ^ succ (k : Ordinal)) + · have := mul_le_mul_left' (one_le_iff_pos.2 <| Nat.cast_pos'.2 n.pos) (ω0 ^ succ (k : Ordinal)) rw [opow_mul] simpa [-opow_succ] · cases m From 81e65433ee271e679931bc2674586fce8fd5db8e Mon Sep 17 00:00:00 2001 From: FR Date: Sat, 19 Oct 2024 05:19:36 +0000 Subject: [PATCH 172/505] chore(Data/Nat/{Bits, Bitwise}): use `Bool.toNat` (#17925) --- Mathlib/Data/Nat/Bits.lean | 6 +++--- Mathlib/Data/Nat/Bitwise.lean | 12 ++++-------- Mathlib/Data/Nat/Size.lean | 4 ++-- 3 files changed, 9 insertions(+), 13 deletions(-) diff --git a/Mathlib/Data/Nat/Bits.lean b/Mathlib/Data/Nat/Bits.lean index 371661d05995b..71f01660fc34f 100644 --- a/Mathlib/Data/Nat/Bits.lean +++ b/Mathlib/Data/Nat/Bits.lean @@ -73,7 +73,7 @@ lemma bodd_mul (m n : ℕ) : bodd (m * n) = (bodd m && bodd n) := by simp only [mul_succ, bodd_add, IH, bodd_succ] cases bodd m <;> cases bodd n <;> rfl -lemma mod_two_of_bodd (n : ℕ) : n % 2 = cond (bodd n) 1 0 := by +lemma mod_two_of_bodd (n : ℕ) : n % 2 = (bodd n).toNat := by have := congr_arg bodd (mod_add_div n 2) simp? [not] at this says simp only [bodd_add, bodd_mul, bodd_succ, not, bodd_zero, Bool.false_and, Bool.bne_false] @@ -100,7 +100,7 @@ lemma div2_succ (n : ℕ) : div2 (n + 1) = cond (bodd n) (succ (div2 n)) (div2 n attribute [local simp] Nat.add_comm Nat.add_assoc Nat.add_left_comm Nat.mul_comm Nat.mul_assoc -lemma bodd_add_div2 : ∀ n, cond (bodd n) 1 0 + 2 * div2 n = n +lemma bodd_add_div2 : ∀ n, (bodd n).toNat + 2 * div2 n = n | 0 => rfl | succ n => by simp only [bodd_succ, Bool.cond_not, div2_succ, Nat.mul_comm] @@ -117,7 +117,7 @@ lemma div2_val (n) : div2 n = n / 2 := by /-- `bit b` appends the digit `b` to the binary representation of its natural number input. -/ def bit (b : Bool) : ℕ → ℕ := cond b (2 * · + 1) (2 * ·) -lemma bit_val (b n) : bit b n = 2 * n + cond b 1 0 := by +lemma bit_val (b n) : bit b n = 2 * n + b.toNat := by cases b <;> rfl lemma bit_decomp (n : Nat) : bit (bodd n) (div2 n) = n := diff --git a/Mathlib/Data/Nat/Bitwise.lean b/Mathlib/Data/Nat/Bitwise.lean index 2a45dc564a4b8..cce15c0d55eee 100644 --- a/Mathlib/Data/Nat/Bitwise.lean +++ b/Mathlib/Data/Nat/Bitwise.lean @@ -89,22 +89,18 @@ lemma bitwise_bit {f : Bool → Bool → Bool} (h : f false false = false := by <;> simp_all (config := {decide := true}) [two_mul] lemma bit_mod_two (a : Bool) (x : ℕ) : - bit a x % 2 = if a then 1 else 0 := by - #adaptation_note /-- nightly-2024-03-16: simp was - -- simp (config := { unfoldPartialApp := true }) only [bit, bit1, bit0, ← mul_two, - -- Bool.cond_eq_ite] -/ - simp only [bit, ite_apply, ← mul_two, Bool.cond_eq_ite] - split_ifs <;> simp [Nat.add_mod] + bit a x % 2 = a.toNat := by + cases a <;> simp [bit_val, mul_add_mod] @[simp] lemma bit_mod_two_eq_zero_iff (a x) : bit a x % 2 = 0 ↔ !a := by - rw [bit_mod_two]; split_ifs <;> simp_all + simp [bit_mod_two] @[simp] lemma bit_mod_two_eq_one_iff (a x) : bit a x % 2 = 1 ↔ a := by - rw [bit_mod_two]; split_ifs <;> simp_all + simp [bit_mod_two] @[simp] theorem lor_bit : ∀ a m b n, bit a m ||| bit b n = bit (a || b) (m ||| n) := diff --git a/Mathlib/Data/Nat/Size.lean b/Mathlib/Data/Nat/Size.lean index 1fdafff0b8fd0..fec3892450d06 100644 --- a/Mathlib/Data/Nat/Size.lean +++ b/Mathlib/Data/Nat/Size.lean @@ -19,8 +19,8 @@ theorem shiftLeft_eq_mul_pow (m) : ∀ n, m <<< n = m * 2 ^ n := shiftLeft_eq _ theorem shiftLeft'_tt_eq_mul_pow (m) : ∀ n, shiftLeft' true m n + 1 = (m + 1) * 2 ^ n | 0 => by simp [shiftLeft', pow_zero, Nat.one_mul] | k + 1 => by - rw [shiftLeft', bit_val, cond_true, add_assoc, ← Nat.mul_add_one, shiftLeft'_tt_eq_mul_pow m k, - mul_left_comm, mul_comm 2, pow_succ] + rw [shiftLeft', bit_val, Bool.toNat_true, add_assoc, ← Nat.mul_add_one, + shiftLeft'_tt_eq_mul_pow m k, mul_left_comm, mul_comm 2, pow_succ] end From 74600feef71163f17723cd491710cbe1cec142d7 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sat, 19 Oct 2024 06:27:29 +0000 Subject: [PATCH 173/505] refactor(Order/Filter): don't use `generate` for `CompleteLattice` (#17799) --- Mathlib/Analysis/Complex/Basic.lean | 4 +- .../RotationNumber/TranslationNumber.lean | 2 +- Mathlib/MeasureTheory/Group/Measure.lean | 2 +- Mathlib/Order/Filter/Basic.lean | 44 ++++++-------- Mathlib/Order/Filter/Defs.lean | 59 +++++++++++-------- 5 files changed, 56 insertions(+), 55 deletions(-) diff --git a/Mathlib/Analysis/Complex/Basic.lean b/Mathlib/Analysis/Complex/Basic.lean index e4257988bc45b..55726f263e5bf 100644 --- a/Mathlib/Analysis/Complex/Basic.lean +++ b/Mathlib/Analysis/Complex/Basic.lean @@ -598,10 +598,10 @@ theorem ofReal_tsum (f : α → ℝ) : (↑(∑' a, f a) : ℂ) = ∑' a, ↑(f RCLike.ofReal_tsum _ _ theorem hasSum_re {f : α → ℂ} {x : ℂ} (h : HasSum f x) : HasSum (fun x => (f x).re) x.re := - RCLike.hasSum_re _ h + RCLike.hasSum_re ℂ h theorem hasSum_im {f : α → ℂ} {x : ℂ} (h : HasSum f x) : HasSum (fun x => (f x).im) x.im := - RCLike.hasSum_im _ h + RCLike.hasSum_im ℂ h theorem re_tsum {f : α → ℂ} (h : Summable f) : (∑' a, f a).re = ∑' a, (f a).re := RCLike.re_tsum _ h diff --git a/Mathlib/Dynamics/Circle/RotationNumber/TranslationNumber.lean b/Mathlib/Dynamics/Circle/RotationNumber/TranslationNumber.lean index 745facc0985ce..65ee8350a4451 100644 --- a/Mathlib/Dynamics/Circle/RotationNumber/TranslationNumber.lean +++ b/Mathlib/Dynamics/Circle/RotationNumber/TranslationNumber.lean @@ -607,7 +607,7 @@ theorem tendsto_translationNumber_of_dist_bounded_aux (x : ℕ → ℝ) (C : ℝ theorem translationNumber_eq_of_dist_bounded {f g : CircleDeg1Lift} (C : ℝ) (H : ∀ n : ℕ, dist ((f ^ n) 0) ((g ^ n) 0) ≤ C) : τ f = τ g := Eq.symm <| g.translationNumber_eq_of_tendsto_aux <| - f.tendsto_translationNumber_of_dist_bounded_aux _ C H + f.tendsto_translationNumber_of_dist_bounded_aux (fun n ↦ (g ^ n) 0) C H @[simp] theorem translationNumber_one : τ 1 = 0 := diff --git a/Mathlib/MeasureTheory/Group/Measure.lean b/Mathlib/MeasureTheory/Group/Measure.lean index 3e36cb114970b..e058bb5e8e9af 100644 --- a/Mathlib/MeasureTheory/Group/Measure.lean +++ b/Mathlib/MeasureTheory/Group/Measure.lean @@ -775,7 +775,7 @@ nonrec theorem _root_.MulEquiv.isHaarMeasure_map [BorelSpace G] [TopologicalGrou [TopologicalGroup H] (e : G ≃* H) (he : Continuous e) (hesymm : Continuous e.symm) : IsHaarMeasure (Measure.map e μ) := let f : G ≃ₜ H := .mk e - isHaarMeasure_map μ e he e.surjective f.closedEmbedding.tendsto_cocompact + isHaarMeasure_map μ (e : G →* H) he e.surjective f.closedEmbedding.tendsto_cocompact /-- A convenience wrapper for MeasureTheory.Measure.isAddHaarMeasure_map`. -/ instance _root_.ContinuousLinearEquiv.isAddHaarMeasure_map diff --git a/Mathlib/Order/Filter/Basic.lean b/Mathlib/Order/Filter/Basic.lean index 84c4397e07ac6..8b546db30b150 100644 --- a/Mathlib/Order/Filter/Basic.lean +++ b/Mathlib/Order/Filter/Basic.lean @@ -83,6 +83,9 @@ instance inhabitedMem : Inhabited { s : Set α // s ∈ f } := theorem filter_eq_iff : f = g ↔ f.sets = g.sets := ⟨congr_arg _, filter_eq⟩ +@[simp] theorem sets_subset_sets : f.sets ⊆ g.sets ↔ g ≤ f := .rfl +@[simp] theorem sets_ssubset_sets : f.sets ⊂ g.sets ↔ g < f := .rfl + /-- An extensionality lemma that is useful for filters with good lemmas about `sᶜ ∈ f` (e.g., `Filter.comap`, `Filter.coprod`, `Filter.Coprod`, `Filter.cofinite`). -/ protected theorem coext (h : ∀ s, sᶜ ∈ f ↔ sᶜ ∈ g) : f = g := @@ -253,21 +256,20 @@ theorem mem_inf_iff_superset {f g : Filter α} {s : Set α} : section CompleteLattice -/- We lift the complete lattice along the Galois connection `generate` / `sets`. Unfortunately, - we want to have different definitional equalities for some lattice operations. So we define them - upfront and change the lattice operations for the complete lattice instance. -/ -instance instCompleteLatticeFilter : CompleteLattice (Filter α) := - { @OrderDual.instCompleteLattice _ (giGenerate α).liftCompleteLattice with - le := (· ≤ ·) - top := ⊤ - le_top := fun _ _s hs => (mem_top.1 hs).symm ▸ univ_mem - inf := (· ⊓ ·) - inf_le_left := fun _ _ _ => mem_inf_of_left - inf_le_right := fun _ _ _ => mem_inf_of_right - le_inf := fun _ _ _ h₁ h₂ _s ⟨_a, ha, _b, hb, hs⟩ => hs.symm ▸ inter_mem (h₁ ha) (h₂ hb) - sSup := join ∘ 𝓟 - le_sSup := fun _ _f hf _s hs => hs hf - sSup_le := fun _ _f hf _s hs _g hg => hf _ hg hs } +/- Complete lattice structure on `Filter α`. -/ +instance instCompleteLatticeFilter : CompleteLattice (Filter α) where + le_sup_left _ _ _ h := h.1 + le_sup_right _ _ _ h := h.2 + sup_le _ _ _ h₁ h₂ _ h := ⟨h₁ h, h₂ h⟩ + inf_le_left _ _ _ := mem_inf_of_left + inf_le_right _ _ _ := mem_inf_of_right + le_inf := fun _ _ _ h₁ h₂ _s ⟨_a, ha, _b, hb, hs⟩ => hs.symm ▸ inter_mem (h₁ ha) (h₂ hb) + le_sSup _ _ h₁ _ h₂ := h₂ h₁ + sSup_le _ _ h₁ _ h₂ _ h₃ := h₁ _ h₃ h₂ + sInf_le _ _ h₁ _ h₂ := by rw [← Filter.sSup_lowerBounds]; exact fun _ h₃ ↦ h₃ h₁ h₂ + le_sInf _ _ h₁ _ h₂ := by rw [← Filter.sSup_lowerBounds] at h₂; exact h₂ h₁ + le_top _ _ := univ_mem' + bot_le _ _ _ := trivial instance : Inhabited (Filter α) := ⟨⊥⟩ @@ -324,10 +326,6 @@ theorem mem_sup {f g : Filter α} {s : Set α} : s ∈ f ⊔ g ↔ s ∈ f ∧ s theorem union_mem_sup {f g : Filter α} {s t : Set α} (hs : s ∈ f) (ht : t ∈ g) : s ∪ t ∈ f ⊔ g := ⟨mem_of_superset hs subset_union_left, mem_of_superset ht subset_union_right⟩ -@[simp] -theorem mem_sSup {x : Set α} {s : Set (Filter α)} : x ∈ sSup s ↔ ∀ f ∈ s, x ∈ (f : Filter α) := - Iff.rfl - @[simp] theorem mem_iSup {x : Set α} {f : ι → Filter α} : x ∈ iSup f ↔ ∀ i, x ∈ f i := by simp only [← Filter.mem_sets, iSup_sets_eq, mem_iInter] @@ -337,7 +335,7 @@ theorem iSup_neBot {f : ι → Filter α} : (⨆ i, f i).NeBot ↔ ∃ i, (f i). simp [neBot_iff] theorem iInf_eq_generate (s : ι → Filter α) : iInf s = generate (⋃ i, (s i).sets) := - show generate _ = generate _ from congr_arg _ <| congr_arg sSup <| (range_comp _ _).symm + eq_of_forall_le_iff fun _ ↦ by simp [le_generate_iff] theorem mem_iInf_of_mem {f : ι → Filter α} (i : ι) {s} (hs : s ∈ f i) : s ∈ ⨅ i, f i := iInf_le f i hs @@ -598,7 +596,7 @@ abbrev coframeMinimalAxioms : Coframe.MinimalAxioms (Filter α) := iInf_sup_le_sup_sInf := fun f s t ⟨h₁, h₂⟩ => by classical rw [iInf_subtype'] - rw [sInf_eq_iInf', iInf_sets_eq_finite, mem_iUnion] at h₂ + rw [sInf_eq_iInf', ← Filter.mem_sets, iInf_sets_eq_finite, mem_iUnion] at h₂ obtain ⟨u, hu⟩ := h₂ rw [← Finset.inf_eq_iInf] at hu suffices ⨅ i : s, f ⊔ ↑i ≤ f ⊔ u.inf fun i => ↑i from this ⟨h₁, hu⟩ @@ -1572,10 +1570,6 @@ instance : LawfulFunctor (Filter : Type u → Type u) where theorem pure_sets (a : α) : (pure a : Filter α).sets = { s | a ∈ s } := rfl -@[simp] -theorem mem_pure {a : α} {s : Set α} : s ∈ (pure a : Filter α) ↔ a ∈ s := - Iff.rfl - @[simp] theorem eventually_pure {a : α} {p : α → Prop} : (∀ᶠ x in pure a, p x) ↔ p a := Iff.rfl diff --git a/Mathlib/Order/Filter/Defs.lean b/Mathlib/Order/Filter/Defs.lean index cd23371bb3f2f..06e6369f2f301 100644 --- a/Mathlib/Order/Filter/Defs.lean +++ b/Mathlib/Order/Filter/Defs.lean @@ -5,6 +5,7 @@ Authors: Johannes Hölzl, Jeremy Avigad -/ import Mathlib.Data.Set.Basic import Mathlib.Order.SetNotation +import Mathlib.Order.Bounds.Defs /-! # Definitions about filters @@ -141,6 +142,15 @@ scoped notation "𝓟" => Filter.principal @[simp] theorem mem_principal : s ∈ 𝓟 t ↔ t ⊆ s := Iff.rfl +/-- `pure x` is the set of sets that contain `x`. It is equal to `𝓟 {x}` but +with this definition we have `s ∈ pure a` defeq `a ∈ s`. -/ +instance : Pure Filter where + pure x := .copy (𝓟 {x}) {s | x ∈ s} fun _ ↦ by simp + +@[simp] +theorem mem_pure {a : α} {s : Set α} : s ∈ (pure a : Filter α) ↔ a ∈ s := + Iff.rfl + /-- The *kernel* of a filter is the intersection of all its sets. -/ def ker (f : Filter α) : Set α := ⋂₀ f.sets @@ -164,11 +174,25 @@ instance : PartialOrder (Filter α) where theorem le_def : f ≤ g ↔ ∀ x ∈ g, x ∈ f := Iff.rfl -instance : Top (Filter α) := - ⟨{ sets := { s | ∀ x, x ∈ s } - univ_sets := fun x => mem_univ x - sets_of_superset := fun hx hxy a => hxy (hx a) - inter_sets := fun hx hy _ => mem_inter (hx _) (hy _) }⟩ +instance instSupSet : SupSet (Filter α) where + sSup S := join (𝓟 S) + +@[simp] theorem mem_sSup {S : Set (Filter α)} : s ∈ sSup S ↔ ∀ f ∈ S, s ∈ f := .rfl + +/-- Infimum of a set of filters. +This definition is marked as irreducible +so that Lean doesn't try to unfold it when unifying expressions. -/ +@[irreducible] +protected def sInf (s : Set (Filter α)) : Filter α := sSup (lowerBounds s) + +instance instInfSet : InfSet (Filter α) where + sInf := Filter.sInf + +protected theorem sSup_lowerBounds (s : Set (Filter α)) : sSup (lowerBounds s) = sInf s := by + simp [sInf, Filter.sInf] + +instance : Top (Filter α) where + top := .copy (sSup (Set.range pure)) {s | ∀ x, x ∈ s} <| by simp theorem mem_top_iff_forall {s : Set α} : s ∈ (⊤ : Filter α) ↔ ∀ x, x ∈ s := Iff.rfl @@ -178,11 +202,7 @@ theorem mem_top {s : Set α} : s ∈ (⊤ : Filter α) ↔ s = univ := by rw [mem_top_iff_forall, eq_univ_iff_forall] instance : Bot (Filter α) where - bot := - { sets := univ - univ_sets := trivial - sets_of_superset := fun _ _ ↦ trivial - inter_sets := fun _ _ ↦ trivial } + bot := .copy (sSup ∅) univ <| by simp @[simp] theorem mem_bot {s : Set α} : s ∈ (⊥ : Filter α) := @@ -190,7 +210,7 @@ theorem mem_bot {s : Set α} : s ∈ (⊥ : Filter α) := /-- The infimum of filters is the filter generated by intersections of elements of the two filters. -/ -instance : Inf (Filter α) := +instance instInf : Inf (Filter α) := ⟨fun f g : Filter α => { sets := { s | ∃ a ∈ f, ∃ b ∈ g, s = a ∩ b } univ_sets := ⟨_, univ_mem, _, univ_mem, by simp⟩ @@ -205,12 +225,8 @@ instance : Inf (Filter α) := ac_rfl }⟩ /-- The supremum of two filters is the filter that contains sets that belong to both filters. -/ -instance : Sup (Filter α) where - sup f g := - { sets := {s | s ∈ f ∧ s ∈ g} - univ_sets := ⟨univ_mem, univ_mem⟩ - sets_of_superset := fun h₁ h₂ ↦ ⟨mem_of_superset h₁.1 h₂, mem_of_superset h₁.2 h₂⟩ - inter_sets := fun h₁ h₂ ↦ ⟨inter_mem h₁.1 h₂.1, inter_mem h₁.2 h₂.2⟩ } +instance instSup : Sup (Filter α) where + sup f g := .copy (sSup {f, g}) {s | s ∈ f ∧ s ∈ g} <| by simp /-- A filter is `NeBot` if it is not equal to `⊥`, or equivalently the empty set does not belong to the filter. Bourbaki include this assumption in the definition of a filter but we prefer to have a @@ -319,15 +335,6 @@ in adding quantifiers to the middle of `Tendsto`s. See def curry (f : Filter α) (g : Filter β) : Filter (α × β) := bind f fun a ↦ map (a, ·) g -/-- `pure x` is the set of sets that contain `x`. It is equal to `𝓟 {x}` but -with this definition we have `s ∈ pure a` defeq `a ∈ s`. -/ -instance : Pure Filter := - ⟨fun x => - { sets := { s | x ∈ s } - inter_sets := And.intro - sets_of_superset := fun hs hst => hst hs - univ_sets := trivial }⟩ - instance : Bind Filter := ⟨@Filter.bind⟩ From a0521c9b7e636372cca49d817075d7da0fd4a681 Mon Sep 17 00:00:00 2001 From: FR Date: Sat, 19 Oct 2024 06:53:15 +0000 Subject: [PATCH 174/505] chore: deprecate `Setoid.Rel` (#16260) --- .../Limits/Shapes/SingleObj.lean | 2 +- Mathlib/Data/Quot.lean | 2 + Mathlib/Data/Setoid/Basic.lean | 94 ++++++++++--------- Mathlib/Data/Setoid/Partition.lean | 38 ++++---- Mathlib/GroupTheory/Congruence/Defs.lean | 8 +- Mathlib/GroupTheory/DoubleCoset.lean | 10 +- Mathlib/GroupTheory/GroupAction/Basic.lean | 17 ++-- Mathlib/GroupTheory/GroupAction/ConjAct.lean | 2 +- Mathlib/GroupTheory/GroupAction/Quotient.lean | 9 +- Mathlib/Order/Partition/Finpartition.lean | 4 +- .../RingTheory/AdicCompletion/Algebra.lean | 2 +- Mathlib/RingTheory/Congruence/Basic.lean | 4 +- .../RingTheory/Valuation/ValuationRing.lean | 14 +-- .../Topology/Algebra/ProperAction/Basic.lean | 2 +- Mathlib/Topology/DiscreteQuotient.lean | 16 ++-- Mathlib/Topology/MetricSpace/Contracting.lean | 4 +- Mathlib/Topology/Separation.lean | 2 +- 17 files changed, 118 insertions(+), 112 deletions(-) diff --git a/Mathlib/CategoryTheory/Limits/Shapes/SingleObj.lean b/Mathlib/CategoryTheory/Limits/Shapes/SingleObj.lean index e501871d0f757..a4e8274b8176e 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/SingleObj.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/SingleObj.lean @@ -75,7 +75,7 @@ variable {G : Type v} [Group G] (J : SingleObj G ⥤ Type u) equivalent to the `MulAction.orbitRel` equivalence relation on `J.obj (SingleObj.star G)`. -/ lemma Types.Quot.Rel.iff_orbitRel (x y : J.obj (SingleObj.star G)) : Types.Quot.Rel J ⟨SingleObj.star G, x⟩ ⟨SingleObj.star G, y⟩ - ↔ Setoid.Rel (MulAction.orbitRel G (J.obj (SingleObj.star G))) x y := by + ↔ MulAction.orbitRel G (J.obj (SingleObj.star G)) x y := by have h (g : G) : y = g • x ↔ g • x = y := ⟨symm, symm⟩ conv => rhs; rw [Setoid.comm'] show (∃ g : G, y = g • x) ↔ (∃ g : G, g • x = y) diff --git a/Mathlib/Data/Quot.lean b/Mathlib/Data/Quot.lean index a11fa73582072..2bec860ba95b9 100644 --- a/Mathlib/Data/Quot.lean +++ b/Mathlib/Data/Quot.lean @@ -26,6 +26,8 @@ run_cmd Lean.Elab.Command.liftTermElabM do Lean.Meta.registerCoercion ``Setoid.r (some { numArgs := 2, coercee := 1, type := .coeFun }) +/-- When writing a lemma about `someSetoid x y` (which uses this instance), +call it `someSetoid_apply` not `someSetoid_r`. -/ instance : CoeFun (Setoid α) (fun _ ↦ α → α → Prop) where coe := @Setoid.r _ diff --git a/Mathlib/Data/Setoid/Basic.lean b/Mathlib/Data/Setoid/Basic.lean index c34439b36e1e1..e4b60325f1b21 100644 --- a/Mathlib/Data/Setoid/Basic.lean +++ b/Mathlib/Data/Setoid/Basic.lean @@ -15,9 +15,6 @@ theorems for quotients of arbitrary types. ## Implementation notes -The function `Rel` and lemmas ending in ' make it easier to talk about different -equivalence relations on the same type. - The complete lattice instance for equivalence relations could have been defined by lifting the Galois insertion of equivalence relations on α into binary relations on α, and then using `CompleteLattice.copy` to define a complete lattice instance with more appropriate @@ -39,47 +36,57 @@ attribute [trans] Setoid.trans variable {α : Type*} {β : Type*} /-- A version of `Setoid.r` that takes the equivalence relation as an explicit argument. -/ +@[deprecated (since := "2024-08-29")] def Setoid.Rel (r : Setoid α) : α → α → Prop := @Setoid.r _ r +set_option linter.deprecated false in +@[deprecated (since := "2024-10-09")] instance Setoid.decidableRel (r : Setoid α) [h : DecidableRel r.r] : DecidableRel r.Rel := h +set_option linter.deprecated false in /-- A version of `Quotient.eq'` compatible with `Setoid.Rel`, to make rewriting possible. -/ +@[deprecated Quotient.eq' (since := "2024-10-09")] theorem Quotient.eq_rel {r : Setoid α} {x y} : (Quotient.mk' x : Quotient r) = Quotient.mk' y ↔ r.Rel x y := Quotient.eq namespace Setoid -@[ext] +attribute [ext] ext + +set_option linter.deprecated false in +@[deprecated Setoid.ext (since := "2024-10-09")] theorem ext' {r s : Setoid α} (H : ∀ a b, r.Rel a b ↔ s.Rel a b) : r = s := ext H -theorem ext_iff {r s : Setoid α} : r = s ↔ ∀ a b, r.Rel a b ↔ s.Rel a b := +set_option linter.deprecated false in +@[deprecated Setoid.ext_iff (since := "2024-10-09")] +theorem ext'_iff {r s : Setoid α} : r = s ↔ ∀ a b, r.Rel a b ↔ s.Rel a b := ⟨fun h _ _ => h ▸ Iff.rfl, ext'⟩ /-- Two equivalence relations are equal iff their underlying binary operations are equal. -/ -theorem eq_iff_rel_eq {r₁ r₂ : Setoid α} : r₁ = r₂ ↔ r₁.Rel = r₂.Rel := - ⟨fun h => h ▸ rfl, fun h => Setoid.ext' fun _ _ => h ▸ Iff.rfl⟩ +theorem eq_iff_rel_eq {r₁ r₂ : Setoid α} : r₁ = r₂ ↔ ⇑r₁ = ⇑r₂ := + ⟨fun h => h ▸ rfl, fun h => Setoid.ext fun _ _ => h ▸ Iff.rfl⟩ /-- Defining `≤` for equivalence relations. -/ instance : LE (Setoid α) := - ⟨fun r s => ∀ ⦃x y⦄, r.Rel x y → s.Rel x y⟩ + ⟨fun r s => ∀ ⦃x y⦄, r x y → s x y⟩ -theorem le_def {r s : Setoid α} : r ≤ s ↔ ∀ {x y}, r.Rel x y → s.Rel x y := +theorem le_def {r s : Setoid α} : r ≤ s ↔ ∀ {x y}, r x y → s x y := Iff.rfl @[refl] -theorem refl' (r : Setoid α) (x) : r.Rel x x := r.iseqv.refl x +theorem refl' (r : Setoid α) (x) : r x x := r.iseqv.refl x @[symm] -theorem symm' (r : Setoid α) : ∀ {x y}, r.Rel x y → r.Rel y x := r.iseqv.symm +theorem symm' (r : Setoid α) : ∀ {x y}, r x y → r y x := r.iseqv.symm @[trans] -theorem trans' (r : Setoid α) : ∀ {x y z}, r.Rel x y → r.Rel y z → r.Rel x z := r.iseqv.trans +theorem trans' (r : Setoid α) : ∀ {x y z}, r x y → r y z → r x z := r.iseqv.trans -theorem comm' (s : Setoid α) {x y} : s.Rel x y ↔ s.Rel y x := +theorem comm' (s : Setoid α) {x y} : s x y ↔ s y x := ⟨s.symm', s.symm'⟩ /-- The kernel of a function is an equivalence relation. -/ @@ -89,7 +96,7 @@ def ker (f : α → β) : Setoid α := /-- The kernel of the quotient map induced by an equivalence relation r equals r. -/ @[simp] theorem ker_mk_eq (r : Setoid α) : ker (@Quotient.mk'' _ r) = r := - ext' fun _ _ => Quotient.eq + ext fun _ _ => Quotient.eq theorem ker_apply_mk_out {f : α → β} (a : α) : f (⟦a⟧ : Quotient (Setoid.ker f)).out = f a := @Quotient.mk_out _ (Setoid.ker f) a @@ -98,7 +105,7 @@ theorem ker_apply_mk_out' {f : α → β} (a : α) : f (Quotient.mk _ a : Quotient <| Setoid.ker f).out' = f a := @Quotient.mk_out' _ (Setoid.ker f) a -theorem ker_def {f : α → β} {x y : α} : (ker f).Rel x y ↔ f x = f y := +theorem ker_def {f : α → β} {x y : α} : ker f x y ↔ f x = f y := Iff.rfl /-- Given types `α`, `β`, the product of two equivalence relations `r` on `α` and `s` on `β`: @@ -106,7 +113,7 @@ theorem ker_def {f : α → β} {x y : α} : (ker f).Rel x y ↔ f x = f y := by `r` and `x₂` is related to `y₂` by `s`. -/ protected def prod (r : Setoid α) (s : Setoid β) : Setoid (α × β) where - r x y := r.Rel x.1 y.1 ∧ s.Rel x.2 y.2 + r x y := r x.1 y.1 ∧ s x.2 y.2 iseqv := ⟨fun x => ⟨r.refl' x.1, s.refl' x.2⟩, fun h => ⟨r.symm' h.1, s.symm' h.2⟩, fun h₁ h₂ => ⟨r.trans' h₁.1 h₂.1, s.trans' h₁.2 h₂.2⟩⟩ @@ -157,28 +164,28 @@ noncomputable def piQuotientEquiv {ι : Sort*} {α : ι → Sort*} (r : ∀ i, S /-- The infimum of two equivalence relations. -/ instance : Inf (Setoid α) := ⟨fun r s => - ⟨fun x y => r.Rel x y ∧ s.Rel x y, + ⟨fun x y => r x y ∧ s x y, ⟨fun x => ⟨r.refl' x, s.refl' x⟩, fun h => ⟨r.symm' h.1, s.symm' h.2⟩, fun h1 h2 => ⟨r.trans' h1.1 h2.1, s.trans' h1.2 h2.2⟩⟩⟩⟩ /-- The infimum of 2 equivalence relations r and s is the same relation as the infimum of the underlying binary operations. -/ -theorem inf_def {r s : Setoid α} : (r ⊓ s).Rel = r.Rel ⊓ s.Rel := +theorem inf_def {r s : Setoid α} : ⇑(r ⊓ s) = ⇑r ⊓ ⇑s := rfl -theorem inf_iff_and {r s : Setoid α} {x y} : (r ⊓ s).Rel x y ↔ r.Rel x y ∧ s.Rel x y := +theorem inf_iff_and {r s : Setoid α} {x y} : (r ⊓ s) x y ↔ r x y ∧ s x y := Iff.rfl /-- The infimum of a set of equivalence relations. -/ instance : InfSet (Setoid α) := ⟨fun S => - { r := fun x y => ∀ r ∈ S, r.Rel x y + { r := fun x y => ∀ r ∈ S, r x y iseqv := ⟨fun x r _ => r.refl' x, fun h r hr => r.symm' <| h r hr, fun h1 h2 r hr => r.trans' (h1 r hr) <| h2 r hr⟩ }⟩ /-- The underlying binary operation of the infimum of a set of equivalence relations is the infimum of the set's image under the map to the underlying binary operation. -/ -theorem sInf_def {s : Set (Setoid α)} : (sInf s).Rel = sInf (Rel '' s) := by +theorem sInf_def {s : Set (Setoid α)} : ⇑(sInf s) = sInf ((⇑) '' s) := by ext simp only [sInf_image, iInf_apply, iInf_Prop_eq] rfl @@ -189,7 +196,7 @@ instance : PartialOrder (Setoid α) where le_refl _ _ _ := id le_trans _ _ _ hr hs _ _ h := hs <| hr h lt_iff_le_not_le _ _ := Iff.rfl - le_antisymm _ _ h1 h2 := Setoid.ext' fun _ _ => ⟨fun h => h1 h, fun h => h2 h⟩ + le_antisymm _ _ h1 h2 := Setoid.ext fun _ _ => ⟨fun h => h1 h, fun h => h2 h⟩ /-- The complete lattice of equivalence relations on a type, with bottom element `=` and top element the trivial equivalence relation. -/ @@ -206,20 +213,20 @@ instance completeLattice : CompleteLattice (Setoid α) := bot_le := fun r x _ h => h ▸ r.2.1 x } @[simp] -theorem top_def : (⊤ : Setoid α).Rel = ⊤ := +theorem top_def : ⇑(⊤ : Setoid α) = ⊤ := rfl @[simp] -theorem bot_def : (⊥ : Setoid α).Rel = (· = ·) := +theorem bot_def : ⇑(⊥ : Setoid α) = (· = ·) := rfl -theorem eq_top_iff {s : Setoid α} : s = (⊤ : Setoid α) ↔ ∀ x y : α, s.Rel x y := by +theorem eq_top_iff {s : Setoid α} : s = (⊤ : Setoid α) ↔ ∀ x y : α, s x y := by rw [_root_.eq_top_iff, Setoid.le_def, Setoid.top_def] simp only [Pi.top_apply, Prop.top_eq_true, forall_true_left] lemma sInf_equiv {S : Set (Setoid α)} {x y : α} : letI := sInf S - x ≈ y ↔ ∀ s ∈ S, s.Rel x y := Iff.rfl + x ≈ y ↔ ∀ s ∈ S, s x y := Iff.rfl lemma sInf_iff {S : Set (Setoid α)} {x y : α} : sInf S x y ↔ ∀ s ∈ S, s x y := Iff.rfl @@ -245,7 +252,7 @@ open Relation /-- The inductively defined equivalence closure of a binary relation r is the infimum of the set of all equivalence relations containing r. -/ theorem eqvGen_eq (r : α → α → Prop) : - EqvGen.setoid r = sInf { s : Setoid α | ∀ ⦃x y⦄, r x y → s.Rel x y } := + EqvGen.setoid r = sInf { s : Setoid α | ∀ ⦃x y⦄, r x y → s x y } := le_antisymm (fun _ _ H => EqvGen.rec (fun _ _ h _ hs => hs h) (refl' _) (fun _ _ _ => symm' _) @@ -255,20 +262,20 @@ theorem eqvGen_eq (r : α → α → Prop) : /-- The supremum of two equivalence relations r and s is the equivalence closure of the binary relation `x is related to y by r or s`. -/ theorem sup_eq_eqvGen (r s : Setoid α) : - r ⊔ s = EqvGen.setoid fun x y => r.Rel x y ∨ s.Rel x y := by + r ⊔ s = EqvGen.setoid fun x y => r x y ∨ s x y := by rw [eqvGen_eq] apply congr_arg sInf simp only [le_def, or_imp, ← forall_and] /-- The supremum of 2 equivalence relations r and s is the equivalence closure of the supremum of the underlying binary operations. -/ -theorem sup_def {r s : Setoid α} : r ⊔ s = EqvGen.setoid (r.Rel ⊔ s.Rel) := by +theorem sup_def {r s : Setoid α} : r ⊔ s = EqvGen.setoid (⇑r ⊔ ⇑s) := by rw [sup_eq_eqvGen]; rfl /-- The supremum of a set S of equivalence relations is the equivalence closure of the binary relation `there exists r ∈ S relating x and y`. -/ theorem sSup_eq_eqvGen (S : Set (Setoid α)) : - sSup S = EqvGen.setoid fun x y => ∃ r : Setoid α, r ∈ S ∧ r.Rel x y := by + sSup S = EqvGen.setoid fun x y => ∃ r : Setoid α, r ∈ S ∧ r x y := by rw [eqvGen_eq] apply congr_arg sInf simp only [upperBounds, le_def, and_imp, exists_imp] @@ -277,7 +284,7 @@ theorem sSup_eq_eqvGen (S : Set (Setoid α)) : /-- The supremum of a set of equivalence relations is the equivalence closure of the supremum of the set's image under the map to the underlying binary operation. -/ -theorem sSup_def {s : Set (Setoid α)} : sSup s = EqvGen.setoid (sSup (Rel '' s)) := by +theorem sSup_def {s : Set (Setoid α)} : sSup s = EqvGen.setoid (sSup ((⇑) '' s)) := by rw [sSup_eq_eqvGen, sSup_image] congr with (x y) simp only [iSup_apply, iSup_Prop_eq, exists_prop] @@ -288,13 +295,12 @@ theorem eqvGen_of_setoid (r : Setoid α) : EqvGen.setoid r.r = r := le_antisymm (by rw [eqvGen_eq]; exact sInf_le fun _ _ => id) EqvGen.rel /-- Equivalence closure is idempotent. -/ -@[simp] -theorem eqvGen_idem (r : α → α → Prop) : EqvGen.setoid (EqvGen.setoid r).Rel = EqvGen.setoid r := +theorem eqvGen_idem (r : α → α → Prop) : EqvGen.setoid (EqvGen.setoid r) = EqvGen.setoid r := eqvGen_of_setoid _ /-- The equivalence closure of a binary relation r is contained in any equivalence relation containing r. -/ -theorem eqvGen_le {r : α → α → Prop} {s : Setoid α} (h : ∀ x y, r x y → s.Rel x y) : +theorem eqvGen_le {r : α → α → Prop} {s : Setoid α} (h : ∀ x y, r x y → s x y) : EqvGen.setoid r ≤ s := by rw [eqvGen_eq]; exact sInf_le h /-- Equivalence closure of binary relations is monotone. -/ @@ -304,7 +310,7 @@ theorem eqvGen_mono {r s : α → α → Prop} (h : ∀ x y, r x y → s x y) : /-- There is a Galois insertion of equivalence relations on α into binary relations on α, with equivalence closure the lower adjoint. -/ -def gi : @GaloisInsertion (α → α → Prop) (Setoid α) _ _ EqvGen.setoid Rel where +def gi : @GaloisInsertion (α → α → Prop) (Setoid α) _ _ EqvGen.setoid (⇑) where choice r _ := EqvGen.setoid r gc _ s := ⟨fun H _ _ h => H <| EqvGen.rel _ _ h, fun H => eqvGen_of_setoid s ▸ eqvGen_mono H⟩ le_l_u x := (eqvGen_of_setoid x).symm ▸ le_refl x @@ -320,7 +326,7 @@ theorem injective_iff_ker_bot (f : α → β) : Injective f ↔ ker f = ⊥ := (@eq_bot_iff (Setoid α) _ _ (ker f)).symm /-- The elements related to x ∈ α by the kernel of f are those in the preimage of f(x) under f. -/ -theorem ker_iff_mem_preimage {f : α → β} {x y} : (ker f).Rel x y ↔ x ∈ f ⁻¹' {f y} := +theorem ker_iff_mem_preimage {f : α → β} {x y} : ker f x y ↔ x ∈ f ⁻¹' {f y} := Iff.rfl /-- Equivalence between functions `α → β` such that `r x y → f x = f y` and functions @@ -345,7 +351,7 @@ theorem ker_lift_injective (f : α → β) : Injective (@Quotient.lift _ _ (ker /-- Given a map f from α to β, the kernel of f is the unique equivalence relation on α whose induced map from the quotient of α to β is injective. -/ -theorem ker_eq_lift_of_injective {r : Setoid α} (f : α → β) (H : ∀ x y, r.Rel x y → f x = f y) +theorem ker_eq_lift_of_injective {r : Setoid α} (f : α → β) (H : ∀ x y, r x y → f x = f y) (h : Injective (Quotient.lift f H)) : ker f = r := le_antisymm (fun x y hk => @@ -387,13 +393,13 @@ variable {r f} closure of the relation on `f`'s image defined by '`x ≈ y` iff the elements of `f⁻¹(x)` are related to the elements of `f⁻¹(y)` by `r`.' -/ def map (r : Setoid α) (f : α → β) : Setoid β := - Relation.EqvGen.setoid fun x y => ∃ a b, f a = x ∧ f b = y ∧ r.Rel a b + Relation.EqvGen.setoid fun x y => ∃ a b, f a = x ∧ f b = y ∧ r a b /-- Given a surjective function f whose kernel is contained in an equivalence relation r, the equivalence relation on f's codomain defined by x ≈ y ↔ the elements of f⁻¹(x) are related to the elements of f⁻¹(y) by r. -/ def mapOfSurjective (r) (f : α → β) (h : ker f ≤ r) (hf : Surjective f) : Setoid β := - ⟨fun x y => ∃ a b, f a = x ∧ f b = y ∧ r.Rel a b, + ⟨fun x y => ∃ a b, f a = x ∧ f b = y ∧ r a b, ⟨fun x => let ⟨y, hy⟩ := hf x ⟨y, y, hy, hy, r.refl' y⟩, @@ -411,9 +417,9 @@ relation on `α` defined by '`x ≈ y` iff `f(x)` is related to `f(y)` by `r`'. See note [reducible non-instances]. -/ abbrev comap (f : α → β) (r : Setoid β) : Setoid α := - ⟨r.Rel on f, r.iseqv.comap _⟩ + ⟨r on f, r.iseqv.comap _⟩ -theorem comap_rel (f : α → β) (r : Setoid β) (x y : α) : (comap f r).Rel x y ↔ r.Rel (f x) (f y) := +theorem comap_rel (f : α → β) (r : Setoid β) (x y : α) : comap f r x y ↔ r (f x) (f y) := Iff.rfl /-- Given a map `f : N → M` and an equivalence relation `r` on `β`, the equivalence relation @@ -424,7 +430,8 @@ theorem comap_eq {f : α → β} {r : Setoid β} : comap f r = ker (@Quotient.mk /-- The second isomorphism theorem for sets. -/ noncomputable def comapQuotientEquiv (f : α → β) (r : Setoid β) : Quotient (comap f r) ≃ Set.range (@Quotient.mk'' _ r ∘ f) := - (Quotient.congrRight <| ext_iff.1 comap_eq).trans <| quotientKerEquivRange <| Quotient.mk'' ∘ f + (Quotient.congrRight <| Setoid.ext_iff.1 comap_eq).trans <| quotientKerEquivRange <| + Quotient.mk'' ∘ f variable (r f) @@ -454,7 +461,7 @@ def correspondence (r : Setoid α) : { s // r ≤ s } ≃o Setoid (Quotient r) w (fun h ↦ s.1.trans' (s.1.trans' (s.2 h₁) h) (s.1.symm' (s.2 h₂))), ⟨Quotient.ind s.1.2.1, @fun x y ↦ Quotient.inductionOn₂ x y fun _ _ ↦ s.1.2.2, @fun x y z ↦ Quotient.inductionOn₃ x y z fun _ _ _ ↦ s.1.2.3⟩⟩ - invFun s := ⟨comap Quotient.mk' s, fun x y h => by rw [comap_rel, eq_rel.2 h]⟩ + invFun s := ⟨comap Quotient.mk' s, fun x y h => by rw [comap_rel, Quotient.eq'.2 h]⟩ left_inv _ := rfl right_inv _ := ext fun x y ↦ Quotient.inductionOn₂ x y fun _ _ ↦ Iff.rfl map_rel_iff' := @@ -478,7 +485,6 @@ theorem Quotient.subsingleton_iff {s : Setoid α} : Subsingleton (Quotient s) refine (surjective_quotient_mk' _).forall.trans (forall_congr' fun a => ?_) refine (surjective_quotient_mk' _).forall.trans (forall_congr' fun b => ?_) simp_rw [Prop.top_eq_true, true_implies, Quotient.eq'] - rfl theorem Quot.subsingleton_iff (r : α → α → Prop) : Subsingleton (Quot r) ↔ Relation.EqvGen r = ⊤ := by diff --git a/Mathlib/Data/Setoid/Partition.lean b/Mathlib/Data/Setoid/Partition.lean index b44853cfa153b..27e668bcb2ece 100644 --- a/Mathlib/Data/Setoid/Partition.lean +++ b/Mathlib/Data/Setoid/Partition.lean @@ -53,9 +53,9 @@ def mkClasses (c : Set (Set α)) (H : ∀ a, ∃! b ∈ c, a ∈ b) : Setoid α /-- Makes the equivalence classes of an equivalence relation. -/ def classes (r : Setoid α) : Set (Set α) := - { s | ∃ y, s = { x | r.Rel x y } } + { s | ∃ y, s = { x | r x y } } -theorem mem_classes (r : Setoid α) (y) : { x | r.Rel x y } ∈ r.classes := +theorem mem_classes (r : Setoid α) (y) : { x | r x y } ∈ r.classes := ⟨y, rfl⟩ theorem classes_ker_subset_fiber_set {β : Type*} (f : α → β) : @@ -74,17 +74,17 @@ theorem card_classes_ker_le {α β : Type*} [Fintype β] (f : α → β) /-- Two equivalence relations are equal iff all their equivalence classes are equal. -/ theorem eq_iff_classes_eq {r₁ r₂ : Setoid α} : - r₁ = r₂ ↔ ∀ x, { y | r₁.Rel x y } = { y | r₂.Rel x y } := - ⟨fun h _x => h ▸ rfl, fun h => ext' fun x => Set.ext_iff.1 <| h x⟩ + r₁ = r₂ ↔ ∀ x, { y | r₁ x y } = { y | r₂ x y } := + ⟨fun h _x => h ▸ rfl, fun h => ext fun x => Set.ext_iff.1 <| h x⟩ -theorem rel_iff_exists_classes (r : Setoid α) {x y} : r.Rel x y ↔ ∃ c ∈ r.classes, x ∈ c ∧ y ∈ c := +theorem rel_iff_exists_classes (r : Setoid α) {x y} : r x y ↔ ∃ c ∈ r.classes, x ∈ c ∧ y ∈ c := ⟨fun h => ⟨_, r.mem_classes y, h, r.refl' y⟩, fun ⟨c, ⟨z, hz⟩, hx, hy⟩ => by subst c exact r.trans' hx (r.symm' hy)⟩ /-- Two equivalence relations are equal iff their equivalence classes are equal. -/ theorem classes_inj {r₁ r₂ : Setoid α} : r₁ = r₂ ↔ r₁.classes = r₂.classes := - ⟨fun h => h ▸ rfl, fun h => ext' fun a b => by simp only [rel_iff_exists_classes, exists_prop, h]⟩ + ⟨fun h => h ▸ rfl, fun h => ext fun a b => by simp only [rel_iff_exists_classes, exists_prop, h]⟩ /-- The empty set is not an equivalence class. -/ theorem empty_not_mem_classes {r : Setoid α} : ∅ ∉ r.classes := fun ⟨y, hy⟩ => @@ -92,7 +92,7 @@ theorem empty_not_mem_classes {r : Setoid α} : ∅ ∉ r.classes := fun ⟨y, h /-- Equivalence classes partition the type. -/ theorem classes_eqv_classes {r : Setoid α} (a) : ∃! b ∈ r.classes, a ∈ b := - ExistsUnique.intro { x | r.Rel x a } ⟨r.mem_classes a, r.refl' _⟩ <| by + ExistsUnique.intro { x | r x a } ⟨r.mem_classes a, r.refl' _⟩ <| by rintro y ⟨⟨_, rfl⟩, ha⟩ ext x exact ⟨fun hx => r.trans' hx (r.symm' ha), fun hx => r.trans' hx ha⟩ @@ -105,7 +105,7 @@ theorem eq_of_mem_classes {r : Setoid α} {x b} (hc : b ∈ r.classes) (hb : x /-- The elements of a set of sets partitioning α are the equivalence classes of the equivalence relation defined by the set of sets. -/ theorem eq_eqv_class_of_mem {c : Set (Set α)} (H : ∀ a, ∃! b ∈ c, a ∈ b) {s y} - (hs : s ∈ c) (hy : y ∈ s) : s = { x | (mkClasses c H).Rel x y } := by + (hs : s ∈ c) (hy : y ∈ s) : s = { x | mkClasses c H x y } := by ext x constructor · intro hx _s' hs' hx' @@ -117,11 +117,11 @@ theorem eq_eqv_class_of_mem {c : Set (Set α)} (H : ∀ a, ∃! b ∈ c, a ∈ b /-- The equivalence classes of the equivalence relation defined by a set of sets partitioning α are elements of the set of sets. -/ theorem eqv_class_mem {c : Set (Set α)} (H : ∀ a, ∃! b ∈ c, a ∈ b) {y} : - { x | (mkClasses c H).Rel x y } ∈ c := + { x | mkClasses c H x y } ∈ c := (H y).elim fun _ hc _ => eq_eqv_class_of_mem H hc.1 hc.2 ▸ hc.1 theorem eqv_class_mem' {c : Set (Set α)} (H : ∀ a, ∃! b ∈ c, a ∈ b) {x} : - { y : α | (mkClasses c H).Rel x y } ∈ c := by + { y : α | mkClasses c H x y } ∈ c := by convert @Setoid.eqv_class_mem _ _ H x using 3 rw [Setoid.comm'] @@ -145,13 +145,13 @@ def setoidOfDisjointUnion {c : Set (Set α)} (hu : Set.sUnion c = @Set.univ α) /-- The equivalence relation made from the equivalence classes of an equivalence relation r equals r. -/ theorem mkClasses_classes (r : Setoid α) : mkClasses r.classes classes_eqv_classes = r := - ext' fun x _y => - ⟨fun h => r.symm' (h { z | r.Rel z x } (r.mem_classes x) <| r.refl' x), fun h _b hb hx => + ext fun x _y => + ⟨fun h => r.symm' (h { z | r z x } (r.mem_classes x) <| r.refl' x), fun h _b hb hx => eq_of_mem_classes (r.mem_classes x) (r.refl' x) hb hx ▸ r.symm' h⟩ @[simp] theorem sUnion_classes (r : Setoid α) : ⋃₀ r.classes = Set.univ := - Set.eq_univ_of_forall fun x => Set.mem_sUnion.2 ⟨{ y | r.Rel y x }, ⟨x, rfl⟩, Setoid.refl _⟩ + Set.eq_univ_of_forall fun x => Set.mem_sUnion.2 ⟨{ y | r y x }, ⟨x, rfl⟩, Setoid.refl _⟩ /-- The equivalence between the quotient by an equivalence relation and its type of equivalence classes. -/ @@ -177,8 +177,8 @@ noncomputable def quotientEquivClasses (r : Setoid α) : Quotient r ≃ Setoid.c @[simp] lemma quotientEquivClasses_mk_eq (r : Setoid α) (a : α) : - (quotientEquivClasses r (Quotient.mk r a) : Set α) = { x | r.Rel x a } := - (@Subtype.ext_iff_val _ _ _ ⟨{ x | r.Rel x a }, Setoid.mem_classes r a⟩).mp rfl + (quotientEquivClasses r (Quotient.mk r a) : Set α) = { x | r x a } := + (@Subtype.ext_iff_val _ _ _ ⟨{ x | r x a }, Setoid.mem_classes r a⟩).mp rfl section Partition @@ -217,7 +217,7 @@ theorem IsPartition.sUnion_eq_univ {c : Set (Set α)} (hc : IsPartition c) : ⋃ /-- All elements of a partition of α are the equivalence class of some y ∈ α. -/ theorem exists_of_mem_partition {c : Set (Set α)} (hc : IsPartition c) {s} (hs : s ∈ c) : - ∃ y, s = { x | (mkClasses c hc.2).Rel x y } := + ∃ y, s = { x | mkClasses c hc.2 x y } := let ⟨y, hy⟩ := nonempty_of_mem_partition hc hs ⟨y, eq_eqv_class_of_mem hc.2 hs hy⟩ @@ -367,7 +367,7 @@ protected abbrev setoid (hs : IndexedPartition s) : Setoid α := theorem index_some (i : ι) : hs.index (hs.some i) = i := (mem_iff_index_eq _).1 <| hs.some_mem i -theorem some_index (x : α) : hs.setoid.Rel (hs.some (hs.index x)) x := +theorem some_index (x : α) : hs.setoid (hs.some (hs.index x)) x := hs.index_some (hs.index x) /-- The quotient associated to an indexed partition. -/ @@ -382,7 +382,7 @@ instance [Inhabited α] : Inhabited hs.Quotient := ⟨hs.proj default⟩ theorem proj_eq_iff {x y : α} : hs.proj x = hs.proj y ↔ hs.index x = hs.index y := - Quotient.eq_rel + Quotient.eq'' @[simp] theorem proj_some_index (x : α) : hs.proj (hs.some (hs.index x)) = hs.proj x := @@ -423,7 +423,7 @@ theorem index_out' (x : hs.Quotient) : hs.index x.out' = hs.index (hs.out x) := theorem proj_out (x : hs.Quotient) : hs.proj (hs.out x) = x := Quotient.inductionOn' x fun x => Quotient.sound' <| hs.some_index x -theorem class_of {x : α} : setOf (hs.setoid.Rel x) = s (hs.index x) := +theorem class_of {x : α} : setOf (hs.setoid x) = s (hs.index x) := Set.ext fun _y => eq_comm.trans hs.mem_iff_index_eq.symm theorem proj_fiber (x : hs.Quotient) : hs.proj ⁻¹' {x} = s (hs.equivQuotient.symm x) := diff --git a/Mathlib/GroupTheory/Congruence/Defs.lean b/Mathlib/GroupTheory/Congruence/Defs.lean index ced9c0f1660d1..118baa87520bd 100644 --- a/Mathlib/GroupTheory/Congruence/Defs.lean +++ b/Mathlib/GroupTheory/Congruence/Defs.lean @@ -166,7 +166,7 @@ theorem ext {c d : Con M} (H : ∀ x y, c x y ↔ d x y) : c = d := @[to_additive "The map sending an additive congruence relation to its underlying equivalence relation is injective."] theorem toSetoid_inj {c d : Con M} (H : c.toSetoid = d.toSetoid) : c = d := - ext <| ext_iff.1 H + ext <| Setoid.ext_iff.1 H /-- Two congruence relations are equal iff their underlying binary relations are equal. -/ @[to_additive "Two additive congruence relations are equal iff their underlying binary relations @@ -334,7 +334,7 @@ instance : InfSet (Con M) where @[to_additive "The infimum of a set of additive congruence relations is the same as the infimum of the set's image under the map to the underlying equivalence relation."] theorem sInf_toSetoid (S : Set (Con M)) : (sInf S).toSetoid = sInf (toSetoid '' S) := - Setoid.ext' fun x y => + Setoid.ext fun x y => ⟨fun h r ⟨c, hS, hr⟩ => by rw [← hr]; exact h c hS, fun h c hS => h c.toSetoid ⟨c, hS, rfl⟩⟩ /-- The infimum of a set of congruence relations is the same as the infimum of the set's image @@ -572,8 +572,8 @@ def correspondence : { d // c ≤ d } ≃o Con c.Quotient where constructor · intros h x y hs rcases h ⟨x, y, rfl, rfl, hs⟩ with ⟨a, b, hx, hy, ht⟩ - exact t.1.trans (t.1.symm <| t.2 <| Quotient.eq_rel.1 hx) - (t.1.trans ht (t.2 <| Quotient.eq_rel.1 hy)) + exact t.1.trans (t.1.symm <| t.2 <| Quotient.eq'.1 hx) + (t.1.trans ht (t.2 <| Quotient.eq'.1 hy)) · intros h _ _ hs rcases hs with ⟨a, b, hx, hy, Hs⟩ exact ⟨a, b, hx, hy, h Hs⟩ diff --git a/Mathlib/GroupTheory/DoubleCoset.lean b/Mathlib/GroupTheory/DoubleCoset.lean index a7d7e29638510..d80688cff9e27 100644 --- a/Mathlib/GroupTheory/DoubleCoset.lean +++ b/Mathlib/GroupTheory/DoubleCoset.lean @@ -72,15 +72,15 @@ def Quotient (H K : Set G) : Type _ := _root_.Quotient (setoid H K) theorem rel_iff {H K : Subgroup G} {x y : G} : - (setoid ↑H ↑K).Rel x y ↔ ∃ a ∈ H, ∃ b ∈ K, y = a * x * b := + setoid ↑H ↑K x y ↔ ∃ a ∈ H, ∃ b ∈ K, y = a * x * b := Iff.trans ⟨fun (hxy : doset x H K = doset y H K) => hxy ▸ mem_doset_self H K y, fun hxy => (doset_eq_of_mem hxy).symm⟩ mem_doset theorem bot_rel_eq_leftRel (H : Subgroup G) : - (setoid ↑(⊥ : Subgroup G) ↑H).Rel = (QuotientGroup.leftRel H).Rel := by + ⇑(setoid ↑(⊥ : Subgroup G) ↑H) = ⇑(QuotientGroup.leftRel H) := by ext a b - rw [rel_iff, Setoid.Rel, QuotientGroup.leftRel_apply] + rw [rel_iff, QuotientGroup.leftRel_apply] constructor · rintro ⟨a, rfl : a = 1, b, hb, rfl⟩ change a⁻¹ * (1 * a * b) ∈ H @@ -89,9 +89,9 @@ theorem bot_rel_eq_leftRel (H : Subgroup G) : exact ⟨1, rfl, a⁻¹ * b, h, by rw [one_mul, mul_inv_cancel_left]⟩ theorem rel_bot_eq_right_group_rel (H : Subgroup G) : - (setoid ↑H ↑(⊥ : Subgroup G)).Rel = (QuotientGroup.rightRel H).Rel := by + ⇑(setoid ↑H ↑(⊥ : Subgroup G)) = ⇑(QuotientGroup.rightRel H) := by ext a b - rw [rel_iff, Setoid.Rel, QuotientGroup.rightRel_apply] + rw [rel_iff, QuotientGroup.rightRel_apply] constructor · rintro ⟨b, hb, a, rfl : a = 1, rfl⟩ change b * a * 1 * a⁻¹ ∈ H diff --git a/Mathlib/GroupTheory/GroupAction/Basic.lean b/Mathlib/GroupTheory/GroupAction/Basic.lean index 9f99144a7769a..547d1fc8f93c9 100644 --- a/Mathlib/GroupTheory/GroupAction/Basic.lean +++ b/Mathlib/GroupTheory/GroupAction/Basic.lean @@ -373,12 +373,11 @@ def orbitRel : Setoid α where variable {G α} @[to_additive] -theorem orbitRel_apply {a b : α} : (orbitRel G α).Rel a b ↔ a ∈ orbit G b := +theorem orbitRel_apply {a b : α} : orbitRel G α a b ↔ a ∈ orbit G b := Iff.rfl -@[to_additive] -lemma orbitRel_r_apply {a b : α} : (orbitRel G _).r a b ↔ a ∈ orbit G b := - Iff.rfl +@[to_additive (attr := deprecated (since := "2024-10-18"))] +alias orbitRel_r_apply := orbitRel_apply @[to_additive] lemma orbitRel_subgroup_le (H : Subgroup G) : orbitRel H α ≤ orbitRel G α := @@ -468,7 +467,7 @@ theorem pretransitive_iff_subsingleton_quotient : · refine Quot.inductionOn a (fun x ↦ ?_) exact Quot.inductionOn b (fun y ↦ Quot.sound <| exists_smul_eq G y x) · have h : Quotient.mk (orbitRel G α) b = ⟦a⟧ := Subsingleton.elim _ _ - exact Quotient.eq_rel.mp h + exact Quotient.eq''.mp h /-- If `α` is non-empty, an action is pretransitive if and only if the quotient has exactly one element. -/ @@ -511,7 +510,7 @@ lemma orbitRel.Quotient.orbit_injective : Injective (orbitRel.Quotient.orbit : orbitRel.Quotient G α → Set α) := by intro x y h simp_rw [orbitRel.Quotient.orbit_eq_orbit_out _ Quotient.out_eq', orbit_eq_iff, - ← orbitRel_r_apply] at h + ← orbitRel_apply] at h simpa [← Quotient.eq''] using h @[to_additive (attr := simp)] @@ -596,7 +595,7 @@ lemma orbitRel.Quotient.mem_subgroup_orbit_iff' {H : Subgroup G} {x : orbitRel.Q at hb rw [orbitRel.Quotient.mem_subgroup_orbit_iff] convert hb using 1 - rw [orbit_eq_iff, ← orbitRel_r_apply, ← Quotient.eq'', Quotient.out_eq', @Quotient.mk''_eq_mk] + rw [orbit_eq_iff, ← orbitRel_apply, ← Quotient.eq'', Quotient.out_eq', @Quotient.mk''_eq_mk] rw [orbitRel.Quotient.mem_orbit, h, @Quotient.mk''_eq_mk] variable (G) (α) @@ -725,7 +724,7 @@ theorem stabilizer_smul_eq_stabilizer_map_conj (g : G) (a : α) : inv_mul_cancel, one_smul, ← mem_stabilizer_iff, Subgroup.mem_map_equiv, MulAut.conj_symm_apply] /-- A bijection between the stabilizers of two elements in the same orbit. -/ -noncomputable def stabilizerEquivStabilizerOfOrbitRel {a b : α} (h : (orbitRel G α).Rel a b) : +noncomputable def stabilizerEquivStabilizerOfOrbitRel {a b : α} (h : orbitRel G α a b) : stabilizer G a ≃* stabilizer G b := let g : G := Classical.choose h have hg : g • b = a := Classical.choose_spec h @@ -749,7 +748,7 @@ theorem stabilizer_vadd_eq_stabilizer_map_conj (g : G) (a : α) : AddAut.conj_symm_apply] /-- A bijection between the stabilizers of two elements in the same orbit. -/ -noncomputable def stabilizerEquivStabilizerOfOrbitRel {a b : α} (h : (orbitRel G α).Rel a b) : +noncomputable def stabilizerEquivStabilizerOfOrbitRel {a b : α} (h : orbitRel G α a b) : stabilizer G a ≃+ stabilizer G b := let g : G := Classical.choose h have hg : g +ᵥ b = a := Classical.choose_spec h diff --git a/Mathlib/GroupTheory/GroupAction/ConjAct.lean b/Mathlib/GroupTheory/GroupAction/ConjAct.lean index 1f5c561c568ef..866e1f3b36304 100644 --- a/Mathlib/GroupTheory/GroupAction/ConjAct.lean +++ b/Mathlib/GroupTheory/GroupAction/ConjAct.lean @@ -246,7 +246,7 @@ theorem fixedPoints_eq_center : fixedPoints (ConjAct G) G = center G := by theorem mem_orbit_conjAct {g h : G} : g ∈ orbit (ConjAct G) h ↔ IsConj g h := by rw [isConj_comm, isConj_iff, mem_orbit_iff]; rfl -theorem orbitRel_conjAct : (orbitRel (ConjAct G) G).Rel = IsConj := +theorem orbitRel_conjAct : ⇑(orbitRel (ConjAct G) G) = IsConj := funext₂ fun g h => by rw [orbitRel_apply, mem_orbit_conjAct] theorem orbit_eq_carrier_conjClasses (g : G) : diff --git a/Mathlib/GroupTheory/GroupAction/Quotient.lean b/Mathlib/GroupTheory/GroupAction/Quotient.lean index 3babc32ec4071..f58416df124ac 100644 --- a/Mathlib/GroupTheory/GroupAction/Quotient.lean +++ b/Mathlib/GroupTheory/GroupAction/Quotient.lean @@ -329,8 +329,7 @@ noncomputable def equivSubgroupOrbitsSetoidComap (H : Subgroup α) (ω : Ω) : have hx := x.property simp only [Set.mem_preimage, Set.mem_singleton_iff] at hx rwa [orbitRel.Quotient.mem_orbit, @Quotient.mk''_eq_mk]⟩⟧) fun a b h ↦ by - change Setoid.Rel _ _ _ at h - rw [Setoid.comap_rel, Setoid.Rel, ← Quotient.eq'', @Quotient.mk''_eq_mk] at h + rw [Setoid.comap_rel, ← Quotient.eq'', @Quotient.mk''_eq_mk] at h simp only [orbitRel.Quotient.subgroup_quotient_eq_iff] exact h left_inv := by @@ -373,7 +372,7 @@ noncomputable def equivSubgroupOrbitsQuotientGroup [IsPretransitive α β] orbitRel.Quotient H β ≃ α ⧸ H where toFun := fun q ↦ q.liftOn' (fun y ↦ (exists_smul_eq α y x).choose) (by intro y₁ y₂ h - rw [orbitRel_r_apply] at h + rw [orbitRel_apply] at h rw [Quotient.eq'', leftRel_eq] dsimp only rcases h with ⟨g, rfl⟩ @@ -389,12 +388,12 @@ noncomputable def equivSubgroupOrbitsQuotientGroup [IsPretransitive α β] intro g₁ g₂ h rw [leftRel_eq] at h simp only - rw [← @Quotient.mk''_eq_mk, Quotient.eq'', orbitRel_r_apply] + rw [← @Quotient.mk''_eq_mk, Quotient.eq'', orbitRel_apply] exact ⟨⟨_, h⟩, by simp [mul_smul]⟩) left_inv := fun y ↦ by induction' y using Quotient.inductionOn' with y simp only [Quotient.liftOn'_mk''] - rw [← @Quotient.mk''_eq_mk, Quotient.eq'', orbitRel_r_apply] + rw [← @Quotient.mk''_eq_mk, Quotient.eq'', orbitRel_apply] convert mem_orbit_self _ rw [inv_smul_eq_iff, (exists_smul_eq α y x).choose_spec] right_inv := fun g ↦ by diff --git a/Mathlib/Order/Partition/Finpartition.lean b/Mathlib/Order/Partition/Finpartition.lean index 2392533c815c4..0c1bbdebaec82 100644 --- a/Mathlib/Order/Partition/Finpartition.lean +++ b/Mathlib/Order/Partition/Finpartition.lean @@ -581,12 +581,12 @@ def ofSetoid (s : Setoid α) [DecidableRel s.r] : Finpartition (univ : Finset α sup_parts := by ext a simp only [sup_image, Function.id_comp, mem_univ, mem_sup, mem_filter, true_and, iff_true] - use a; exact s.refl a + use a not_bot_mem := by rw [bot_eq_empty, mem_image, not_exists] intro a simp only [filter_eq_empty_iff, not_forall, mem_univ, forall_true_left, true_and, not_not] - use a; exact s.refl a + use a theorem mem_part_ofSetoid_iff_rel {s : Setoid α} [DecidableRel s.r] {b : α} : b ∈ (ofSetoid s).part a ↔ s.r a b := by diff --git a/Mathlib/RingTheory/AdicCompletion/Algebra.lean b/Mathlib/RingTheory/AdicCompletion/Algebra.lean index f9ab3648fed20..6f7c6168e2c6c 100644 --- a/Mathlib/RingTheory/AdicCompletion/Algebra.lean +++ b/Mathlib/RingTheory/AdicCompletion/Algebra.lean @@ -197,7 +197,7 @@ theorem smul_mk {m n : ℕ} (hmn : m ≤ n) (r : AdicCauchySequence I R) good definitional behaviour for the module instance on adic completions -/ instance : SMul (R ⧸ (I • ⊤ : Ideal R)) (M ⧸ (I • ⊤ : Submodule R M)) where smul r x := - Quotient.liftOn r (· • x) fun b₁ b₂ (h : Setoid.Rel _ b₁ b₂) ↦ by + Quotient.liftOn r (· • x) fun b₁ b₂ h ↦ by refine Quotient.inductionOn' x (fun x ↦ ?_) have h : b₁ - b₂ ∈ (I : Submodule R R) := by rwa [show I = I • ⊤ by simp, ← Submodule.quotientRel_def] diff --git a/Mathlib/RingTheory/Congruence/Basic.lean b/Mathlib/RingTheory/Congruence/Basic.lean index 131575077c37b..90eea0cd93488 100644 --- a/Mathlib/RingTheory/Congruence/Basic.lean +++ b/Mathlib/RingTheory/Congruence/Basic.lean @@ -74,7 +74,7 @@ instance : FunLike (RingCon R) R (R → Prop) where rcases x with ⟨⟨x, _⟩, _⟩ rcases y with ⟨⟨y, _⟩, _⟩ congr! - rw [Setoid.ext_iff, (show x.Rel = y.Rel from h)] + rw [Setoid.ext_iff, (show ⇑x = ⇑y from h)] simp theorem rel_eq_coe : c.r = c := @@ -445,7 +445,7 @@ instance : InfSet (RingCon R) where /-- The infimum of a set of congruence relations is the same as the infimum of the set's image under the map to the underlying equivalence relation. -/ theorem sInf_toSetoid (S : Set (RingCon R)) : (sInf S).toSetoid = sInf ((·.toSetoid) '' S) := - Setoid.ext' fun x y => + Setoid.ext fun x y => ⟨fun h r ⟨c, hS, hr⟩ => by rw [← hr]; exact h c hS, fun h c hS => h c.toSetoid ⟨c, hS, rfl⟩⟩ /-- The infimum of a set of congruence relations is the same as the infimum of the set's image diff --git a/Mathlib/RingTheory/Valuation/ValuationRing.lean b/Mathlib/RingTheory/Valuation/ValuationRing.lean index e95d7d322ff0d..266cd8e9d5884 100644 --- a/Mathlib/RingTheory/Valuation/ValuationRing.lean +++ b/Mathlib/RingTheory/Valuation/ValuationRing.lean @@ -142,22 +142,22 @@ noncomputable instance linearOrder : LinearOrder (ValueGroup A K) where noncomputable instance linearOrderedCommGroupWithZero : LinearOrderedCommGroupWithZero (ValueGroup A K) := { linearOrder .. with - mul_assoc := by rintro ⟨a⟩ ⟨b⟩ ⟨c⟩; apply Quotient.sound'; rw [mul_assoc]; apply Setoid.refl' - one_mul := by rintro ⟨a⟩; apply Quotient.sound'; rw [one_mul]; apply Setoid.refl' - mul_one := by rintro ⟨a⟩; apply Quotient.sound'; rw [mul_one]; apply Setoid.refl' - mul_comm := by rintro ⟨a⟩ ⟨b⟩; apply Quotient.sound'; rw [mul_comm]; apply Setoid.refl' + mul_assoc := by rintro ⟨a⟩ ⟨b⟩ ⟨c⟩; apply Quotient.sound'; rw [mul_assoc] + one_mul := by rintro ⟨a⟩; apply Quotient.sound'; rw [one_mul] + mul_one := by rintro ⟨a⟩; apply Quotient.sound'; rw [mul_one] + mul_comm := by rintro ⟨a⟩ ⟨b⟩; apply Quotient.sound'; rw [mul_comm] mul_le_mul_left := by rintro ⟨a⟩ ⟨b⟩ ⟨c, rfl⟩ ⟨d⟩ use c; simp only [Algebra.smul_def]; ring - zero_mul := by rintro ⟨a⟩; apply Quotient.sound'; rw [zero_mul]; apply Setoid.refl' - mul_zero := by rintro ⟨a⟩; apply Quotient.sound'; rw [mul_zero]; apply Setoid.refl' + zero_mul := by rintro ⟨a⟩; apply Quotient.sound'; rw [zero_mul] + mul_zero := by rintro ⟨a⟩; apply Quotient.sound'; rw [mul_zero] zero_le_one := ⟨0, by rw [zero_smul]⟩ exists_pair_ne := by use 0, 1 intro c; obtain ⟨d, hd⟩ := Quotient.exact' c apply_fun fun t => d⁻¹ • t at hd simp only [inv_smul_smul, smul_zero, one_ne_zero] at hd - inv_zero := by apply Quotient.sound'; rw [inv_zero]; apply Setoid.refl' + inv_zero := by apply Quotient.sound'; rw [inv_zero] mul_inv_cancel := by rintro ⟨a⟩ ha apply Quotient.sound' diff --git a/Mathlib/Topology/Algebra/ProperAction/Basic.lean b/Mathlib/Topology/Algebra/ProperAction/Basic.lean index fa7a99510f216..92fc69d56d9f4 100644 --- a/Mathlib/Topology/Algebra/ProperAction/Basic.lean +++ b/Mathlib/Topology/Algebra/ProperAction/Basic.lean @@ -118,7 +118,7 @@ theorem t2Space_quotient_mulAction_of_properSMul [ProperSMul G X] : · ext ⟨x₁, x₂⟩ simp only [mem_preimage, map_apply, mem_diagonal_iff, mem_range, Prod.mk.injEq, Prod.exists, exists_eq_right] - rw [Quotient.eq_rel, MulAction.orbitRel_apply, MulAction.mem_orbit_iff] + rw [Quotient.eq', MulAction.orbitRel_apply, MulAction.mem_orbit_iff] all_goals infer_instance /-- If a T2 group acts properly on a topological space, then this topological space is T2. -/ diff --git a/Mathlib/Topology/DiscreteQuotient.lean b/Mathlib/Topology/DiscreteQuotient.lean index ed3253cad0993..bad892ec52aeb 100644 --- a/Mathlib/Topology/DiscreteQuotient.lean +++ b/Mathlib/Topology/DiscreteQuotient.lean @@ -69,7 +69,7 @@ variable {α X Y Z : Type*} [TopologicalSpace X] [TopologicalSpace Y] [Topologic @[ext] -- Porting note: in Lean 4, uses projection to `r` instead of `Setoid`. structure DiscreteQuotient (X : Type*) [TopologicalSpace X] extends Setoid X where /-- For every point `x`, the set `{ y | Rel x y }` is an open set. -/ - protected isOpen_setOf_rel : ∀ x, IsOpen (setOf (toSetoid.Rel x)) + protected isOpen_setOf_rel : ∀ x, IsOpen (setOf (toSetoid x)) namespace DiscreteQuotient @@ -81,13 +81,13 @@ lemma toSetoid_injective : Function.Injective (@toSetoid X _) /-- Construct a discrete quotient from a clopen set. -/ def ofIsClopen {A : Set X} (h : IsClopen A) : DiscreteQuotient X where toSetoid := ⟨fun x y => x ∈ A ↔ y ∈ A, fun _ => Iff.rfl, Iff.symm, Iff.trans⟩ - isOpen_setOf_rel x := by by_cases hx : x ∈ A <;> simp [Setoid.Rel, hx, h.1, h.2, ← compl_setOf] + isOpen_setOf_rel x := by by_cases hx : x ∈ A <;> simp [hx, h.1, h.2, ← compl_setOf] -theorem refl : ∀ x, S.Rel x x := S.refl' +theorem refl : ∀ x, S.toSetoid x x := S.refl' -theorem symm (x y : X) : S.Rel x y → S.Rel y x := S.symm' +theorem symm (x y : X) : S.toSetoid x y → S.toSetoid y x := S.symm' -theorem trans (x y z : X) : S.Rel x y → S.Rel y z → S.Rel x z := S.trans' +theorem trans (x y z : X) : S.toSetoid x y → S.toSetoid y z → S.toSetoid x z := S.trans' /-- The setoid whose quotient yields the discrete quotient. -/ add_decl_doc toSetoid @@ -101,7 +101,7 @@ instance : TopologicalSpace S := /-- The projection from `X` to the given discrete quotient. -/ def proj : X → S := Quotient.mk'' -theorem fiber_eq (x : X) : S.proj ⁻¹' {S.proj x} = setOf (S.Rel x) := +theorem fiber_eq (x : X) : S.proj ⁻¹' {S.proj x} = setOf (S.toSetoid x) := Set.ext fun _ => eq_comm.trans Quotient.eq'' theorem proj_surjective : Function.Surjective S.proj := @@ -130,7 +130,7 @@ theorem isOpen_preimage (A : Set S) : IsOpen (S.proj ⁻¹' A) := theorem isClosed_preimage (A : Set S) : IsClosed (S.proj ⁻¹' A) := (S.isClopen_preimage A).1 -theorem isClopen_setOf_rel (x : X) : IsClopen (setOf (S.Rel x)) := by +theorem isClopen_setOf_rel (x : X) : IsClopen (setOf (S.toSetoid x)) := by rw [← fiber_eq] apply isClopen_preimage @@ -359,7 +359,7 @@ lemma comp_finsetClopens [CompactSpace X] : (Set.image (fun (t : Clopens X) ↦ t.carrier) ∘ Finset.toSet) ∘ finsetClopens X = fun ⟨f, _⟩ ↦ f.classes := by ext d - simp only [Setoid.classes, Setoid.Rel, Set.mem_setOf_eq, Function.comp_apply, + simp only [Setoid.classes, Set.mem_setOf_eq, Function.comp_apply, finsetClopens, Set.coe_toFinset, Set.mem_image, Set.mem_range, exists_exists_eq_and] constructor diff --git a/Mathlib/Topology/MetricSpace/Contracting.lean b/Mathlib/Topology/MetricSpace/Contracting.lean index 2fcec8e182a86..ac3f24a9253fa 100644 --- a/Mathlib/Topology/MetricSpace/Contracting.lean +++ b/Mathlib/Topology/MetricSpace/Contracting.lean @@ -137,7 +137,7 @@ theorem efixedPoint_eq_of_edist_lt_top (hf : ContractingWith K f) {x : α} (hx : efixedPoint f hf x hx = efixedPoint f hf y hy := by refine (hf.eq_or_edist_eq_top_of_fixedPoints ?_ ?_).elim id fun h' ↦ False.elim (ne_of_lt ?_ h') <;> try apply efixedPoint_isFixedPt - change edistLtTopSetoid.Rel _ _ + change edistLtTopSetoid _ _ trans x · apply Setoid.symm' -- Porting note: Originally `symm` exact hf.edist_efixedPoint_lt_top hx @@ -221,7 +221,7 @@ theorem efixedPoint_eq_of_edist_lt_top' (hf : ContractingWith K f) {s : Set α} efixedPoint' f hsc hsf hfs x hxs hx = efixedPoint' f htc htf hft y hyt hy := by refine (hf.eq_or_edist_eq_top_of_fixedPoints ?_ ?_).elim id fun h' ↦ False.elim (ne_of_lt ?_ h') <;> try apply efixedPoint_isFixedPt' - change edistLtTopSetoid.Rel _ _ + change edistLtTopSetoid _ _ trans x · apply Setoid.symm' -- Porting note: Originally `symm` apply edist_efixedPoint_lt_top' diff --git a/Mathlib/Topology/Separation.lean b/Mathlib/Topology/Separation.lean index 1ba06927812eb..d2e132acfed38 100644 --- a/Mathlib/Topology/Separation.lean +++ b/Mathlib/Topology/Separation.lean @@ -1632,7 +1632,7 @@ instance : TopologicalSpace (t2Quotient X) := /-- The map from a topological space to its largest T2 quotient. -/ def mk : X → t2Quotient X := Quotient.mk (t2Setoid X) -lemma mk_eq {x y : X} : mk x = mk y ↔ ∀ s : Setoid X, T2Space (Quotient s) → s.Rel x y := +lemma mk_eq {x y : X} : mk x = mk y ↔ ∀ s : Setoid X, T2Space (Quotient s) → s x y := Setoid.quotient_mk_sInf_eq variable (X) From 90c6be64b3540bf42afd057554434af4597f1367 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 19 Oct 2024 06:53:17 +0000 Subject: [PATCH 175/505] =?UTF-8?q?feat(MeasureTheory):=20`=CE=BC.comap=20?= =?UTF-8?q?Prod.swap=20=3D=20=CE=BC.map=20Prod.swap`=20(#17918)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit From PFR --- Mathlib/MeasureTheory/Measure/MeasureSpace.lean | 7 +++++++ 1 file changed, 7 insertions(+) diff --git a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean index 7a7ab5f8c32c6..3c03fb33601d8 100644 --- a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean +++ b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean @@ -1975,6 +1975,13 @@ theorem quasiMeasurePreserving_symm (μ : Measure α) (e : α ≃ᵐ β) : end MeasurableEquiv +namespace MeasureTheory.Measure +variable {mα : MeasurableSpace α} {mβ : MeasurableSpace β} + +lemma comap_swap (μ : Measure (α × β)) : μ.comap Prod.swap = μ.map Prod.swap := + (MeasurableEquiv.prodComm ..).comap_symm + +end MeasureTheory.Measure end set_option linter.style.longFile 2000 From 61509b3d7569a6b81e6686f22dd8db1b56156e09 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Sat, 19 Oct 2024 07:30:32 +0000 Subject: [PATCH 176/505] chore(SetTheory/Ordinal/Basic): make `Ordinal.typein` a `PrincipalSeg` (#17599) This allows us to golf most of its API down to one-liners. --- Mathlib/Order/InitialSeg.lean | 7 ++ Mathlib/SetTheory/Ordinal/Basic.lean | 113 ++++++++++++++------------- 2 files changed, 65 insertions(+), 55 deletions(-) diff --git a/Mathlib/Order/InitialSeg.lean b/Mathlib/Order/InitialSeg.lean index e53d78efa0277..b1e211077acf2 100644 --- a/Mathlib/Order/InitialSeg.lean +++ b/Mathlib/Order/InitialSeg.lean @@ -289,6 +289,10 @@ theorem mem_range_of_rel [IsTrans β s] (f : r ≺i s) {a : α} {b : β} (h : s @[deprecated mem_range_of_rel (since := "2024-09-21")] alias init := mem_range_of_rel +theorem surjOn (f : r ≺i s) : Set.SurjOn f Set.univ { b | s b f.top } := by + intro b h + simpa using mem_range_of_rel_top _ h + /-- A principal segment is in particular an initial segment. -/ instance hasCoeInitialSeg [IsTrans β s] : Coe (r ≺i s) (r ≼i s) := ⟨fun f => ⟨f.toRelEmbedding, fun _ _ => f.mem_range_of_rel⟩⟩ @@ -383,6 +387,9 @@ instance [IsWellOrder β s] : Subsingleton (r ≺i s) := cases g have := RelEmbedding.coe_fn_injective ef; congr ⟩ +protected theorem eq [IsWellOrder β s] (f g : r ≺i s) (a) : f a = g a := by + rw [Subsingleton.elim f g] + theorem top_eq [IsWellOrder γ t] (e : r ≃r s) (f : r ≺i t) (g : s ≺i t) : f.top = g.top := by rw [Subsingleton.elim f (PrincipalSeg.equivLT e g)]; rfl diff --git a/Mathlib/SetTheory/Ordinal/Basic.lean b/Mathlib/SetTheory/Ordinal/Basic.lean index 2de607c6ac406..bc2a81d9ed4ff 100644 --- a/Mathlib/SetTheory/Ordinal/Basic.lean +++ b/Mathlib/SetTheory/Ordinal/Basic.lean @@ -139,11 +139,6 @@ instance inhabited : Inhabited Ordinal := instance one : One Ordinal := ⟨type <| @EmptyRelation PUnit⟩ -/-- The order type of an element inside a well order. For the embedding as a principal segment, see -`typein.principalSeg`. -/ -def typein (r : α → α → Prop) [IsWellOrder α r] (a : α) : Ordinal := - type (Subrel r { b | r b a }) - @[simp] theorem type_def' (w : WellOrder) : ⟦w⟧ = type w.r := by cases w @@ -375,69 +370,73 @@ def principalSegToType {α β : Ordinal} (h : α < β) : α.toType by - rcases f.mem_range_of_rel_top h with ⟨b, rfl⟩; exact ⟨b, rfl⟩⟩ - -@[simp] -theorem typein_apply {α β} {r : α → α → Prop} {s : β → β → Prop} [IsWellOrder α r] [IsWellOrder β s] - (f : r ≼i s) (a : α) : Ordinal.typein s (f a) = Ordinal.typein r a := - Eq.symm <| - Quotient.sound - ⟨RelIso.ofSurjective - (RelEmbedding.codRestrict _ ((Subrel.relEmbedding _ _).trans f) fun ⟨x, h⟩ => by - rw [RelEmbedding.trans_apply]; exact f.toRelEmbedding.map_rel_iff.2 h) - fun ⟨y, h⟩ => by - rcases f.mem_range_of_rel h with ⟨a, rfl⟩ - exact ⟨⟨a, f.toRelEmbedding.map_rel_iff.1 h⟩, - Subtype.eq <| RelEmbedding.trans_apply _ _ _⟩⟩ +theorem typein_top {α β} {r : α → α → Prop} {s : β → β → Prop} + [IsWellOrder α r] [IsWellOrder β s] (f : r ≺i s) : typein s f.top = type r := + f.subrelIso.ordinal_type_eq @[simp] theorem typein_lt_typein (r : α → α → Prop) [IsWellOrder α r] {a b : α} : typein r a < typein r b ↔ r a b := - ⟨fun ⟨f⟩ => by - have : f.top.1 = a := by - let f' := PrincipalSeg.ofElement r a - let g' := f.trans (PrincipalSeg.ofElement r b) - have : g'.top = f'.top := by rw [Subsingleton.elim f' g'] - exact this - rw [← this] - exact f.top.2, fun h => - ⟨PrincipalSeg.codRestrict _ (PrincipalSeg.ofElement r a) (fun x => @trans _ r _ _ _ _ x.2 h) h⟩⟩ + (typein r).map_rel_iff + +theorem mem_range_typein_iff (r : α → α → Prop) [IsWellOrder α r] {o} : + o ∈ Set.range (typein r) ↔ o < type r := + (typein r).mem_range_iff_rel theorem typein_surj (r : α → α → Prop) [IsWellOrder α r] {o} (h : o < type r) : - ∃ a, typein r a = o := - inductionOn o (fun _ _ _ ⟨f⟩ => ⟨f.top, typein_top _⟩) h + o ∈ Set.range (typein r) := + (typein r).mem_range_of_rel_top h + +theorem typein_surjOn (r : α → α → Prop) [IsWellOrder α r] : + Set.SurjOn (typein r) Set.univ (Set.Iio (type r)) := + (typein r).surjOn theorem typein_injective (r : α → α → Prop) [IsWellOrder α r] : Injective (typein r) := - injective_of_increasing r (· < ·) (typein r) (typein_lt_typein r).2 + (typein r).injective -@[simp] theorem typein_inj (r : α → α → Prop) [IsWellOrder α r] {a b} : typein r a = typein r b ↔ a = b := - (typein_injective r).eq_iff - -/-- Principal segment version of the `typein` function, embedding a well order into ordinals as a -principal segment. -/ -def typein.principalSeg {α : Type u} (r : α → α → Prop) [IsWellOrder α r] : - @PrincipalSeg α Ordinal.{u} r (· < ·) := - ⟨⟨⟨typein r, typein_injective r⟩, typein_lt_typein r⟩, type r, - fun _ ↦ ⟨fun ⟨a, h⟩ ↦ h ▸ typein_lt_type r a, typein_surj r⟩⟩ - -@[simp] -theorem typein.principalSeg_coe (r : α → α → Prop) [IsWellOrder α r] : - (typein.principalSeg r : α → Ordinal) = typein r := - rfl + (typein r).inj /-! ### Enumerating elements in a well-order with ordinals. -/ @@ -450,16 +449,16 @@ the elements of `α`. -/ @[simps! symm_apply_coe] def enum (r : α → α → Prop) [IsWellOrder α r] : @RelIso { o // o < type r } α (Subrel (· < ·) { o | o < type r }) r := - (typein.principalSeg r).subrelIso + (typein r).subrelIso @[simp] theorem typein_enum (r : α → α → Prop) [IsWellOrder α r] {o} (h : o < type r) : typein r (enum r ⟨o, h⟩) = o := - (typein.principalSeg r).apply_subrelIso _ + (typein r).apply_subrelIso _ theorem enum_type {α β} {r : α → α → Prop} {s : β → β → Prop} [IsWellOrder α r] [IsWellOrder β s] (f : s ≺i r) {h : type s < type r} : enum r ⟨type s, h⟩ = f.top := - (typein.principalSeg r).injective <| (typein_enum _ _).trans (typein_top _).symm + (typein r).injective <| (typein_enum _ _).trans (typein_top _).symm @[simp] theorem enum_typein (r : α → α → Prop) [IsWellOrder α r] (a : α) : @@ -502,6 +501,10 @@ theorem induction {p : Ordinal.{u} → Prop} (i : Ordinal.{u}) (h : ∀ j, (∀ p i := lt_wf.induction i h +theorem typein_apply {α β} {r : α → α → Prop} {s : β → β → Prop} [IsWellOrder α r] [IsWellOrder β s] + (f : r ≼i s) (a : α) : typein s (f a) = typein r a := by + rw [← f.leLT_apply _ a, (f.leLT _).eq] + /-! ### Cardinality of ordinals -/ @@ -1031,7 +1034,7 @@ def liftPrincipalSeg : Ordinal.{u} Date: Sat, 19 Oct 2024 07:30:33 +0000 Subject: [PATCH 177/505] chore(Algebra/Module): split off `NoZeroSMulDivisors` (#17909) The definition of `NoZeroSMulDivisors` doesn't require us to know about modules. By splitting off this definition, we can streamline the downstream import graph a bit. In the process, generalize a few instances from `Module` to `SMulWithZero`. New files: * `NoZeroSMulDivisors/Defs.lean`: definition and basic results that don't require `Module` * `NoZeroSMulDivisors/Basic.lean`: basic results that require `Module` * `NoZeroSMulDivisors/Pi.lean`: results on `NoZeroSMulDivisors _ (_ -> _)` that don't require `Module` * `NoZeroSMulDivisors/Prod.lean`: results on `NoZeroSMulDivisors _ (_ x _)` that don't require `Module` --- Mathlib.lean | 4 + Mathlib/Algebra/Algebra/Basic.lean | 1 + Mathlib/Algebra/BigOperators/Finprod.lean | 2 +- Mathlib/Algebra/Group/EvenFunction.lean | 4 +- Mathlib/Algebra/ModEq.lean | 1 + Mathlib/Algebra/Module/Basic.lean | 25 +-- Mathlib/Algebra/Module/Card.lean | 2 +- Mathlib/Algebra/Module/Defs.lean | 176 ------------------ Mathlib/Algebra/Module/LinearMap/Basic.lean | 2 +- Mathlib/Algebra/Module/LinearMap/End.lean | 1 + Mathlib/Algebra/Module/Pi.lean | 13 -- Mathlib/Algebra/Module/Prod.lean | 13 -- Mathlib/Algebra/Module/Rat.lean | 1 + Mathlib/Algebra/NoZeroSMulDivisors/Basic.lean | 155 +++++++++++++++ Mathlib/Algebra/NoZeroSMulDivisors/Defs.lean | 78 ++++++++ Mathlib/Algebra/NoZeroSMulDivisors/Pi.lean | 34 ++++ Mathlib/Algebra/NoZeroSMulDivisors/Prod.lean | 32 ++++ Mathlib/Algebra/Order/Module/Defs.lean | 2 +- Mathlib/Algebra/Order/Star/Basic.lean | 1 + Mathlib/Data/ENNReal/Operations.lean | 1 - Mathlib/Data/Set/Pointwise/SMul.lean | 1 + Mathlib/GroupTheory/Subgroup/Saturated.lean | 2 +- Mathlib/LinearAlgebra/BilinearMap.lean | 1 + Mathlib/LinearAlgebra/LinearPMap.lean | 1 - Mathlib/LinearAlgebra/Multilinear/Basic.lean | 1 + Mathlib/LinearAlgebra/Span.lean | 1 + Mathlib/RingTheory/Nilpotent/Basic.lean | 2 +- 27 files changed, 321 insertions(+), 236 deletions(-) create mode 100644 Mathlib/Algebra/NoZeroSMulDivisors/Basic.lean create mode 100644 Mathlib/Algebra/NoZeroSMulDivisors/Defs.lean create mode 100644 Mathlib/Algebra/NoZeroSMulDivisors/Pi.lean create mode 100644 Mathlib/Algebra/NoZeroSMulDivisors/Prod.lean diff --git a/Mathlib.lean b/Mathlib.lean index c66a3419a9aeb..66e48c38b8ae0 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -544,6 +544,10 @@ import Mathlib.Algebra.MvPolynomial.Rename import Mathlib.Algebra.MvPolynomial.Supported import Mathlib.Algebra.MvPolynomial.Variables import Mathlib.Algebra.NeZero +import Mathlib.Algebra.NoZeroSMulDivisors.Basic +import Mathlib.Algebra.NoZeroSMulDivisors.Defs +import Mathlib.Algebra.NoZeroSMulDivisors.Pi +import Mathlib.Algebra.NoZeroSMulDivisors.Prod import Mathlib.Algebra.Notation import Mathlib.Algebra.Opposites import Mathlib.Algebra.Order.AbsoluteValue diff --git a/Mathlib/Algebra/Algebra/Basic.lean b/Mathlib/Algebra/Algebra/Basic.lean index 66932daf4008d..e9e129ad5c0fd 100644 --- a/Mathlib/Algebra/Algebra/Basic.lean +++ b/Mathlib/Algebra/Algebra/Basic.lean @@ -9,6 +9,7 @@ import Mathlib.Algebra.Module.Equiv.Basic import Mathlib.Algebra.Module.Submodule.Ker import Mathlib.Algebra.Module.Submodule.RestrictScalars import Mathlib.Algebra.Module.ULift +import Mathlib.Algebra.NoZeroSMulDivisors.Basic import Mathlib.Algebra.Ring.Subring.Basic import Mathlib.Data.Nat.Cast.Order.Basic import Mathlib.Data.Int.CharZero diff --git a/Mathlib/Algebra/BigOperators/Finprod.lean b/Mathlib/Algebra/BigOperators/Finprod.lean index fbfabaaebfca3..b6b2d96428073 100644 --- a/Mathlib/Algebra/BigOperators/Finprod.lean +++ b/Mathlib/Algebra/BigOperators/Finprod.lean @@ -5,7 +5,7 @@ Authors: Kexing Ying, Kevin Buzzard, Yury Kudryashov -/ import Mathlib.Algebra.BigOperators.GroupWithZero.Finset import Mathlib.Algebra.Group.FiniteSupport -import Mathlib.Algebra.Module.Defs +import Mathlib.Algebra.NoZeroSMulDivisors.Basic import Mathlib.Algebra.Order.BigOperators.Group.Finset import Mathlib.Data.Set.Subsingleton diff --git a/Mathlib/Algebra/Group/EvenFunction.lean b/Mathlib/Algebra/Group/EvenFunction.lean index 4713a0549c61f..81ddf35dcd57c 100644 --- a/Mathlib/Algebra/Group/EvenFunction.lean +++ b/Mathlib/Algebra/Group/EvenFunction.lean @@ -3,9 +3,9 @@ Copyright (c) 2024 David Loeffler. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: David Loeffler -/ -import Mathlib.Algebra.Group.Action.Pi -import Mathlib.Algebra.Module.Defs import Mathlib.Algebra.BigOperators.Group.Finset +import Mathlib.Algebra.Group.Action.Pi +import Mathlib.Algebra.NoZeroSMulDivisors.Basic /-! # Even and odd functions diff --git a/Mathlib/Algebra/ModEq.lean b/Mathlib/Algebra/ModEq.lean index a275aad2c7254..2744ba932b93f 100644 --- a/Mathlib/Algebra/ModEq.lean +++ b/Mathlib/Algebra/ModEq.lean @@ -5,6 +5,7 @@ Authors: Yaël Dillies -/ import Mathlib.Data.Int.ModEq import Mathlib.Algebra.Field.Basic +import Mathlib.Algebra.NoZeroSMulDivisors.Basic import Mathlib.Algebra.Order.Ring.Int import Mathlib.GroupTheory.QuotientGroup.Basic diff --git a/Mathlib/Algebra/Module/Basic.lean b/Mathlib/Algebra/Module/Basic.lean index 8656e47667f24..18deb70317a22 100644 --- a/Mathlib/Algebra/Module/Basic.lean +++ b/Mathlib/Algebra/Module/Basic.lean @@ -7,6 +7,7 @@ import Mathlib.Algebra.Field.Basic import Mathlib.Algebra.Group.Action.Pi import Mathlib.Algebra.Group.Indicator import Mathlib.Algebra.Module.Defs +import Mathlib.Algebra.NoZeroSMulDivisors.Defs /-! # Further basic results about modules. @@ -95,30 +96,6 @@ theorem inv_intCast_smul_comm {α E : Type*} (R : Type*) [AddCommGroup E] [Divis @[deprecated (since := "2024-04-17")] alias inv_int_cast_smul_comm := inv_intCast_smul_comm - - -section NoZeroSMulDivisors - -section Module - -instance [AddCommGroup M] [NoZeroSMulDivisors ℤ M] : NoZeroSMulDivisors ℕ M := - ⟨fun {c x} hcx ↦ by rwa [← Nat.cast_smul_eq_nsmul ℤ c x, smul_eq_zero, Nat.cast_eq_zero] at hcx⟩ - -end Module - -section GroupWithZero - -variable [GroupWithZero R] [AddMonoid M] [DistribMulAction R M] - --- see note [lower instance priority] -/-- This instance applies to `DivisionSemiring`s, in particular `NNReal` and `NNRat`. -/ -instance (priority := 100) GroupWithZero.toNoZeroSMulDivisors : NoZeroSMulDivisors R M := - ⟨fun {a _} h ↦ or_iff_not_imp_left.2 fun ha ↦ (smul_eq_zero_iff_eq <| Units.mk0 a ha).1 h⟩ - -end GroupWithZero - -end NoZeroSMulDivisors - namespace Function lemma support_smul_subset_left [Zero R] [Zero M] [SMulWithZero R M] (f : α → R) (g : α → M) : diff --git a/Mathlib/Algebra/Module/Card.lean b/Mathlib/Algebra/Module/Card.lean index e9c8b1d5446ec..3941ffd972d1d 100644 --- a/Mathlib/Algebra/Module/Card.lean +++ b/Mathlib/Algebra/Module/Card.lean @@ -3,7 +3,7 @@ Copyright (c) 2023 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ -import Mathlib.Algebra.Module.Defs +import Mathlib.Algebra.NoZeroSMulDivisors.Basic import Mathlib.SetTheory.Cardinal.Basic /-! diff --git a/Mathlib/Algebra/Module/Defs.lean b/Mathlib/Algebra/Module/Defs.lean index a024554de986a..c51d03d323fe8 100644 --- a/Mathlib/Algebra/Module/Defs.lean +++ b/Mathlib/Algebra/Module/Defs.lean @@ -408,182 +408,6 @@ instance AddCommGroup.intIsScalarTower {R : Type u} {M : Type v} [Ring R] [AddCo [Module R M] : IsScalarTower ℤ R M where smul_assoc n x y := ((smulAddHom R M).flip y).map_zsmul x n -section NoZeroSMulDivisors - -/-! ### `NoZeroSMulDivisors` - -This section defines the `NoZeroSMulDivisors` class, and includes some tests -for the vanishing of elements (especially in modules over division rings). --/ - - -/-- `NoZeroSMulDivisors R M` states that a scalar multiple is `0` only if either argument is `0`. -This is a version of saying that `M` is torsion free, without assuming `R` is zero-divisor free. - -The main application of `NoZeroSMulDivisors R M`, when `M` is a module, -is the result `smul_eq_zero`: a scalar multiple is `0` iff either argument is `0`. - -It is a generalization of the `NoZeroDivisors` class to heterogeneous multiplication. --/ -@[mk_iff] -class NoZeroSMulDivisors (R M : Type*) [Zero R] [Zero M] [SMul R M] : Prop where - /-- If scalar multiplication yields zero, either the scalar or the vector was zero. -/ - eq_zero_or_eq_zero_of_smul_eq_zero : ∀ {c : R} {x : M}, c • x = 0 → c = 0 ∨ x = 0 - -export NoZeroSMulDivisors (eq_zero_or_eq_zero_of_smul_eq_zero) - -/-- Pullback a `NoZeroSMulDivisors` instance along an injective function. -/ -theorem Function.Injective.noZeroSMulDivisors {R M N : Type*} [Zero R] [Zero M] [Zero N] - [SMul R M] [SMul R N] [NoZeroSMulDivisors R N] (f : M → N) (hf : Function.Injective f) - (h0 : f 0 = 0) (hs : ∀ (c : R) (x : M), f (c • x) = c • f x) : NoZeroSMulDivisors R M := - ⟨fun {_ _} h => - Or.imp_right (@hf _ _) <| h0.symm ▸ eq_zero_or_eq_zero_of_smul_eq_zero (by rw [← hs, h, h0])⟩ - --- See note [lower instance priority] -instance (priority := 100) NoZeroDivisors.toNoZeroSMulDivisors [Zero R] [Mul R] - [NoZeroDivisors R] : NoZeroSMulDivisors R R := - ⟨fun {_ _} => eq_zero_or_eq_zero_of_mul_eq_zero⟩ - -theorem smul_ne_zero [Zero R] [Zero M] [SMul R M] [NoZeroSMulDivisors R M] {c : R} {x : M} - (hc : c ≠ 0) (hx : x ≠ 0) : c • x ≠ 0 := fun h => - (eq_zero_or_eq_zero_of_smul_eq_zero h).elim hc hx - -section SMulWithZero - -variable [Zero R] [Zero M] [SMulWithZero R M] [NoZeroSMulDivisors R M] {c : R} {x : M} - -@[simp] -theorem smul_eq_zero : c • x = 0 ↔ c = 0 ∨ x = 0 := - ⟨eq_zero_or_eq_zero_of_smul_eq_zero, fun h => - h.elim (fun h => h.symm ▸ zero_smul R x) fun h => h.symm ▸ smul_zero c⟩ - -theorem smul_ne_zero_iff : c • x ≠ 0 ↔ c ≠ 0 ∧ x ≠ 0 := by rw [Ne, smul_eq_zero, not_or] - -lemma smul_eq_zero_iff_left (hx : x ≠ 0) : c • x = 0 ↔ c = 0 := by simp [hx] -lemma smul_eq_zero_iff_right (hc : c ≠ 0) : c • x = 0 ↔ x = 0 := by simp [hc] -lemma smul_ne_zero_iff_left (hx : x ≠ 0) : c • x ≠ 0 ↔ c ≠ 0 := by simp [hx] -lemma smul_ne_zero_iff_right (hc : c ≠ 0) : c • x ≠ 0 ↔ x ≠ 0 := by simp [hc] - -end SMulWithZero - -section Module - -section Nat - -theorem Nat.noZeroSMulDivisors - (R) (M) [Semiring R] [CharZero R] [AddCommMonoid M] [Module R M] [NoZeroSMulDivisors R M] : - NoZeroSMulDivisors ℕ M where - eq_zero_or_eq_zero_of_smul_eq_zero {c x} := by rw [← Nat.cast_smul_eq_nsmul R, smul_eq_zero]; simp - -theorem two_nsmul_eq_zero - (R) (M) [Semiring R] [CharZero R] [AddCommMonoid M] [Module R M] [NoZeroSMulDivisors R M] - {v : M} : 2 • v = 0 ↔ v = 0 := by - haveI := Nat.noZeroSMulDivisors R M - simp [smul_eq_zero] - -end Nat - -variable [Semiring R] -variable (R M) - -/-- If `M` is an `R`-module with one and `M` has characteristic zero, then `R` has characteristic -zero as well. Usually `M` is an `R`-algebra. -/ -theorem CharZero.of_module (M) [AddCommMonoidWithOne M] [CharZero M] [Module R M] : CharZero R := by - refine ⟨fun m n h => @Nat.cast_injective M _ _ _ _ ?_⟩ - rw [← nsmul_one, ← nsmul_one, ← Nat.cast_smul_eq_nsmul R, ← Nat.cast_smul_eq_nsmul R, h] - -end Module - -section AddCommGroup - --- `R` can still be a semiring here -variable [Semiring R] [AddCommGroup M] [Module R M] - -section SMulInjective - -variable (M) - -theorem smul_right_injective [NoZeroSMulDivisors R M] {c : R} (hc : c ≠ 0) : - Function.Injective (c • · : M → M) := - (injective_iff_map_eq_zero (smulAddHom R M c)).2 fun _ ha => (smul_eq_zero.mp ha).resolve_left hc - -variable {M} - -theorem smul_right_inj [NoZeroSMulDivisors R M] {c : R} (hc : c ≠ 0) {x y : M} : - c • x = c • y ↔ x = y := - (smul_right_injective M hc).eq_iff - -end SMulInjective - -section Nat - -theorem self_eq_neg - (R) (M) [Semiring R] [CharZero R] [AddCommGroup M] [Module R M] [NoZeroSMulDivisors R M] - {v : M} : v = -v ↔ v = 0 := by - rw [← two_nsmul_eq_zero R M, two_smul, add_eq_zero_iff_eq_neg] - -theorem neg_eq_self - (R) (M) [Semiring R] [CharZero R] [AddCommGroup M] [Module R M] [NoZeroSMulDivisors R M] - {v : M} : -v = v ↔ v = 0 := by - rw [eq_comm, self_eq_neg R M] - -theorem self_ne_neg - (R) (M) [Semiring R] [CharZero R] [AddCommGroup M] [Module R M] [NoZeroSMulDivisors R M] - {v : M} : v ≠ -v ↔ v ≠ 0 := - (self_eq_neg R M).not - -theorem neg_ne_self - (R) (M) [Semiring R] [CharZero R] [AddCommGroup M] [Module R M] [NoZeroSMulDivisors R M] - {v : M} : -v ≠ v ↔ v ≠ 0 := - (neg_eq_self R M).not - -end Nat - -end AddCommGroup - -section Module - -variable [Ring R] [AddCommGroup M] [Module R M] - -section SMulInjective - -variable (R) -variable [NoZeroSMulDivisors R M] - -theorem smul_left_injective {x : M} (hx : x ≠ 0) : Function.Injective fun c : R => c • x := - fun c d h => - sub_eq_zero.mp - ((smul_eq_zero.mp - (calc - (c - d) • x = c • x - d • x := sub_smul c d x - _ = 0 := sub_eq_zero.mpr h - )).resolve_right - hx) - -end SMulInjective - -instance [NoZeroSMulDivisors ℤ M] : NoZeroSMulDivisors ℕ M := - ⟨fun {c x} hcx ↦ by rwa [← Nat.cast_smul_eq_nsmul ℤ, smul_eq_zero, Nat.cast_eq_zero] at hcx⟩ - -variable (R M) - -theorem NoZeroSMulDivisors.int_of_charZero - (R) (M) [Ring R] [AddCommGroup M] [Module R M] [NoZeroSMulDivisors R M] [CharZero R] : - NoZeroSMulDivisors ℤ M := - ⟨fun {z x} h ↦ by simpa [← smul_one_smul R z x] using h⟩ - -/-- Only a ring of characteristic zero can have a non-trivial module without additive or -scalar torsion. -/ -theorem CharZero.of_noZeroSMulDivisors [Nontrivial M] [NoZeroSMulDivisors ℤ M] : CharZero R := by - refine ⟨fun {n m h} ↦ ?_⟩ - obtain ⟨x, hx⟩ := exists_ne (0 : M) - replace h : (n : ℤ) • x = (m : ℤ) • x := by simp [← Nat.cast_smul_eq_nsmul R, h] - simpa using smul_left_injective ℤ hx h - -end Module - -end NoZeroSMulDivisors - -- Porting note (#10618): simp can prove this --@[simp] theorem Nat.smul_one_eq_cast {R : Type*} [NonAssocSemiring R] (m : ℕ) : m • (1 : R) = ↑m := by diff --git a/Mathlib/Algebra/Module/LinearMap/Basic.lean b/Mathlib/Algebra/Module/LinearMap/Basic.lean index 1a68ecbd89506..baff9dbf891d4 100644 --- a/Mathlib/Algebra/Module/LinearMap/Basic.lean +++ b/Mathlib/Algebra/Module/LinearMap/Basic.lean @@ -5,7 +5,7 @@ Authors: Nathaniel Thomas, Jeremy Avigad, Johannes Hölzl, Mario Carneiro, Anne Frédéric Dupuis, Heather Macbeth -/ import Mathlib.Algebra.Module.LinearMap.Defs -import Mathlib.Algebra.Module.Pi +import Mathlib.Algebra.NoZeroSMulDivisors.Pi import Mathlib.GroupTheory.GroupAction.DomAct.Basic /-! diff --git a/Mathlib/Algebra/Module/LinearMap/End.lean b/Mathlib/Algebra/Module/LinearMap/End.lean index f75595188a8c1..a861127390ad8 100644 --- a/Mathlib/Algebra/Module/LinearMap/End.lean +++ b/Mathlib/Algebra/Module/LinearMap/End.lean @@ -6,6 +6,7 @@ Authors: Nathaniel Thomas, Jeremy Avigad, Johannes Hölzl, Mario Carneiro, Anne -/ import Mathlib.Algebra.GroupPower.IterateHom import Mathlib.Algebra.Module.LinearMap.Defs +import Mathlib.Algebra.NoZeroSMulDivisors.Defs /-! # Endomorphisms of a module diff --git a/Mathlib/Algebra/Module/Pi.lean b/Mathlib/Algebra/Module/Pi.lean index 1d6a967b8db91..23c86123bb947 100644 --- a/Mathlib/Algebra/Module/Pi.lean +++ b/Mathlib/Algebra/Module/Pi.lean @@ -83,17 +83,4 @@ instance module' {g : I → Type*} {r : ∀ i, Semiring (f i)} {m : ∀ i, AddCo -- Porting note: not sure why `apply zero_smul` fails here. rw [zero_smul] -instance noZeroSMulDivisors (α) [Semiring α] [∀ i, AddCommMonoid <| f i] - [∀ i, Module α <| f i] [∀ i, NoZeroSMulDivisors α <| f i] : - NoZeroSMulDivisors α (∀ i : I, f i) := - ⟨fun {_ _} h => - or_iff_not_imp_left.mpr fun hc => - funext fun i => (smul_eq_zero.mp (congr_fun h i)).resolve_left hc⟩ - -/-- A special case of `Pi.noZeroSMulDivisors` for non-dependent types. Lean struggles to -synthesize this instance by itself elsewhere in the library. -/ -instance _root_.Function.noZeroSMulDivisors {ι α β : Type*} [Semiring α] [AddCommMonoid β] - [Module α β] [NoZeroSMulDivisors α β] : NoZeroSMulDivisors α (ι → β) := - Pi.noZeroSMulDivisors _ - end Pi diff --git a/Mathlib/Algebra/Module/Prod.lean b/Mathlib/Algebra/Module/Prod.lean index 11b934777a06a..de41fcccad536 100644 --- a/Mathlib/Algebra/Module/Prod.lean +++ b/Mathlib/Algebra/Module/Prod.lean @@ -35,17 +35,4 @@ instance instModule [Semiring R] [AddCommMonoid M] [AddCommMonoid N] [Module R M add_smul := fun _ _ _ => mk.inj_iff.mpr ⟨add_smul _ _ _, add_smul _ _ _⟩ zero_smul := fun _ => mk.inj_iff.mpr ⟨zero_smul _ _, zero_smul _ _⟩ } -instance noZeroSMulDivisors {r : Semiring R} [AddCommMonoid M] [AddCommMonoid N] - [Module R M] [Module R N] [NoZeroSMulDivisors R M] [NoZeroSMulDivisors R N] : - NoZeroSMulDivisors R (M × N) := - { eq_zero_or_eq_zero_of_smul_eq_zero := by -- Porting note: in mathlib3 there is no need for `by`/ - -- `intro`/`exact`, i.e. the following works: - -- ⟨fun c ⟨x, y⟩ h => - -- or_iff_not_imp_left.mpr fun hc => - intro c ⟨x, y⟩ h - exact or_iff_not_imp_left.mpr fun hc => - mk.inj_iff.mpr - ⟨(smul_eq_zero.mp (congr_arg fst h)).resolve_left hc, - (smul_eq_zero.mp (congr_arg snd h)).resolve_left hc⟩ } - end Prod diff --git a/Mathlib/Algebra/Module/Rat.lean b/Mathlib/Algebra/Module/Rat.lean index c0ed84a4366ed..3ce5e5ac011a8 100644 --- a/Mathlib/Algebra/Module/Rat.lean +++ b/Mathlib/Algebra/Module/Rat.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Nathaniel Thomas, Jeremy Avigad, Johannes Hölzl, Mario Carneiro -/ import Mathlib.Algebra.Module.Basic +import Mathlib.Algebra.NoZeroSMulDivisors.Basic import Mathlib.Algebra.Field.Rat import Mathlib.Algebra.Order.Field.Rat diff --git a/Mathlib/Algebra/NoZeroSMulDivisors/Basic.lean b/Mathlib/Algebra/NoZeroSMulDivisors/Basic.lean new file mode 100644 index 0000000000000..f33114e9583a3 --- /dev/null +++ b/Mathlib/Algebra/NoZeroSMulDivisors/Basic.lean @@ -0,0 +1,155 @@ +/- +Copyright (c) 2015 Nathaniel Thomas. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Anne Baanen, Yury Kudryashov, Joseph Myers, Heather Macbeth, Kim Morrison, Yaël Dillies +-/ +import Mathlib.Algebra.Module.Defs +import Mathlib.Algebra.NoZeroSMulDivisors.Defs + +/-! +# `NoZeroSMulDivisors` + +This file defines the `NoZeroSMulDivisors` class, and includes some tests +for the vanishing of elements (especially in modules over division rings). +-/ + +assert_not_exists Multiset +assert_not_exists Set.indicator +assert_not_exists Pi.single_smul₀ +assert_not_exists Field + +section NoZeroSMulDivisors + +variable {R M : Type*} + +section Module + +section Nat + +theorem Nat.noZeroSMulDivisors + (R) (M) [Semiring R] [CharZero R] [AddCommMonoid M] [Module R M] [NoZeroSMulDivisors R M] : + NoZeroSMulDivisors ℕ M where + eq_zero_or_eq_zero_of_smul_eq_zero {c x} := by rw [← Nat.cast_smul_eq_nsmul R, smul_eq_zero]; simp + +theorem two_nsmul_eq_zero + (R) (M) [Semiring R] [CharZero R] [AddCommMonoid M] [Module R M] [NoZeroSMulDivisors R M] + {v : M} : 2 • v = 0 ↔ v = 0 := by + haveI := Nat.noZeroSMulDivisors R M + simp [smul_eq_zero] + +end Nat + +variable [Semiring R] +variable (R M) + +/-- If `M` is an `R`-module with one and `M` has characteristic zero, then `R` has characteristic +zero as well. Usually `M` is an `R`-algebra. -/ +theorem CharZero.of_module (M) [AddCommMonoidWithOne M] [CharZero M] [Module R M] : CharZero R := by + refine ⟨fun m n h => @Nat.cast_injective M _ _ _ _ ?_⟩ + rw [← nsmul_one, ← nsmul_one, ← Nat.cast_smul_eq_nsmul R, ← Nat.cast_smul_eq_nsmul R, h] + +end Module + +section AddCommGroup + +-- `R` can still be a semiring here +variable [Semiring R] [AddCommGroup M] [Module R M] + +section SMulInjective + +variable (M) + +theorem smul_right_injective [NoZeroSMulDivisors R M] {c : R} (hc : c ≠ 0) : + Function.Injective (c • · : M → M) := + (injective_iff_map_eq_zero (smulAddHom R M c)).2 fun _ ha => (smul_eq_zero.mp ha).resolve_left hc + +variable {M} + +theorem smul_right_inj [NoZeroSMulDivisors R M] {c : R} (hc : c ≠ 0) {x y : M} : + c • x = c • y ↔ x = y := + (smul_right_injective M hc).eq_iff + +end SMulInjective + +section Nat + +theorem self_eq_neg + (R) (M) [Semiring R] [CharZero R] [AddCommGroup M] [Module R M] [NoZeroSMulDivisors R M] + {v : M} : v = -v ↔ v = 0 := by + rw [← two_nsmul_eq_zero R M, two_smul, add_eq_zero_iff_eq_neg] + +theorem neg_eq_self + (R) (M) [Semiring R] [CharZero R] [AddCommGroup M] [Module R M] [NoZeroSMulDivisors R M] + {v : M} : -v = v ↔ v = 0 := by + rw [eq_comm, self_eq_neg R M] + +theorem self_ne_neg + (R) (M) [Semiring R] [CharZero R] [AddCommGroup M] [Module R M] [NoZeroSMulDivisors R M] + {v : M} : v ≠ -v ↔ v ≠ 0 := + (self_eq_neg R M).not + +theorem neg_ne_self + (R) (M) [Semiring R] [CharZero R] [AddCommGroup M] [Module R M] [NoZeroSMulDivisors R M] + {v : M} : -v ≠ v ↔ v ≠ 0 := + (neg_eq_self R M).not + +end Nat + +end AddCommGroup + +section Module + +variable [Ring R] [AddCommGroup M] [Module R M] + +section SMulInjective + +variable (R) +variable [NoZeroSMulDivisors R M] + +theorem smul_left_injective {x : M} (hx : x ≠ 0) : Function.Injective fun c : R => c • x := + fun c d h => + sub_eq_zero.mp + ((smul_eq_zero.mp + (calc + (c - d) • x = c • x - d • x := sub_smul c d x + _ = 0 := sub_eq_zero.mpr h + )).resolve_right + hx) + +end SMulInjective + +instance [NoZeroSMulDivisors ℤ M] : NoZeroSMulDivisors ℕ M := + ⟨fun {c x} hcx ↦ by rwa [← Nat.cast_smul_eq_nsmul ℤ, smul_eq_zero, Nat.cast_eq_zero] at hcx⟩ + +variable (R M) + +theorem NoZeroSMulDivisors.int_of_charZero + (R) (M) [Ring R] [AddCommGroup M] [Module R M] [NoZeroSMulDivisors R M] [CharZero R] : + NoZeroSMulDivisors ℤ M := + ⟨fun {z x} h ↦ by simpa [← smul_one_smul R z x] using h⟩ + +/-- Only a ring of characteristic zero can have a non-trivial module without additive or +scalar torsion. -/ +theorem CharZero.of_noZeroSMulDivisors [Nontrivial M] [NoZeroSMulDivisors ℤ M] : CharZero R := by + refine ⟨fun {n m h} ↦ ?_⟩ + obtain ⟨x, hx⟩ := exists_ne (0 : M) + replace h : (n : ℤ) • x = (m : ℤ) • x := by simp [← Nat.cast_smul_eq_nsmul R, h] + simpa using smul_left_injective ℤ hx h + +instance [AddCommGroup M] [NoZeroSMulDivisors ℤ M] : NoZeroSMulDivisors ℕ M := + ⟨fun {c x} hcx ↦ by rwa [← Nat.cast_smul_eq_nsmul ℤ c x, smul_eq_zero, Nat.cast_eq_zero] at hcx⟩ + +end Module + +section GroupWithZero + +variable [GroupWithZero R] [AddMonoid M] [DistribMulAction R M] + +-- see note [lower instance priority] +/-- This instance applies to `DivisionSemiring`s, in particular `NNReal` and `NNRat`. -/ +instance (priority := 100) GroupWithZero.toNoZeroSMulDivisors : NoZeroSMulDivisors R M := + ⟨fun {a _} h ↦ or_iff_not_imp_left.2 fun ha ↦ (smul_eq_zero_iff_eq <| Units.mk0 a ha).1 h⟩ + +end GroupWithZero + +end NoZeroSMulDivisors diff --git a/Mathlib/Algebra/NoZeroSMulDivisors/Defs.lean b/Mathlib/Algebra/NoZeroSMulDivisors/Defs.lean new file mode 100644 index 0000000000000..caa70e20f08e3 --- /dev/null +++ b/Mathlib/Algebra/NoZeroSMulDivisors/Defs.lean @@ -0,0 +1,78 @@ +/- +Copyright (c) 2015 Nathaniel Thomas. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Anne Baanen, Yury Kudryashov, Joseph Myers, Heather Macbeth, Kim Morrison, Yaël Dillies +-/ +import Mathlib.Algebra.SMulWithZero + +/-! +# `NoZeroSMulDivisors` + +This file defines the `NoZeroSMulDivisors` class, and includes some tests +for the vanishing of elements (especially in modules over division rings). +-/ + +assert_not_exists Multiset +assert_not_exists Set.indicator +assert_not_exists Pi.single_smul₀ +assert_not_exists Field +assert_not_exists Module + +section NoZeroSMulDivisors + +/-! ### `NoZeroSMulDivisors` + +-/ + +variable {R M : Type*} + +/-- `NoZeroSMulDivisors R M` states that a scalar multiple is `0` only if either argument is `0`. +This is a version of saying that `M` is torsion free, without assuming `R` is zero-divisor free. + +The main application of `NoZeroSMulDivisors R M`, when `M` is a module, +is the result `smul_eq_zero`: a scalar multiple is `0` iff either argument is `0`. + +It is a generalization of the `NoZeroDivisors` class to heterogeneous multiplication. +-/ +@[mk_iff] +class NoZeroSMulDivisors (R M : Type*) [Zero R] [Zero M] [SMul R M] : Prop where + /-- If scalar multiplication yields zero, either the scalar or the vector was zero. -/ + eq_zero_or_eq_zero_of_smul_eq_zero : ∀ {c : R} {x : M}, c • x = 0 → c = 0 ∨ x = 0 + +export NoZeroSMulDivisors (eq_zero_or_eq_zero_of_smul_eq_zero) + +/-- Pullback a `NoZeroSMulDivisors` instance along an injective function. -/ +theorem Function.Injective.noZeroSMulDivisors {R M N : Type*} [Zero R] [Zero M] [Zero N] + [SMul R M] [SMul R N] [NoZeroSMulDivisors R N] (f : M → N) (hf : Function.Injective f) + (h0 : f 0 = 0) (hs : ∀ (c : R) (x : M), f (c • x) = c • f x) : NoZeroSMulDivisors R M := + ⟨fun {_ _} h => + Or.imp_right (@hf _ _) <| h0.symm ▸ eq_zero_or_eq_zero_of_smul_eq_zero (by rw [← hs, h, h0])⟩ + +-- See note [lower instance priority] +instance (priority := 100) NoZeroDivisors.toNoZeroSMulDivisors [Zero R] [Mul R] + [NoZeroDivisors R] : NoZeroSMulDivisors R R := + ⟨fun {_ _} => eq_zero_or_eq_zero_of_mul_eq_zero⟩ + +theorem smul_ne_zero [Zero R] [Zero M] [SMul R M] [NoZeroSMulDivisors R M] {c : R} {x : M} + (hc : c ≠ 0) (hx : x ≠ 0) : c • x ≠ 0 := fun h => + (eq_zero_or_eq_zero_of_smul_eq_zero h).elim hc hx + +section SMulWithZero + +variable [Zero R] [Zero M] [SMulWithZero R M] [NoZeroSMulDivisors R M] {c : R} {x : M} + +@[simp] +theorem smul_eq_zero : c • x = 0 ↔ c = 0 ∨ x = 0 := + ⟨eq_zero_or_eq_zero_of_smul_eq_zero, fun h => + h.elim (fun h => h.symm ▸ zero_smul R x) fun h => h.symm ▸ smul_zero c⟩ + +theorem smul_ne_zero_iff : c • x ≠ 0 ↔ c ≠ 0 ∧ x ≠ 0 := by rw [Ne, smul_eq_zero, not_or] + +lemma smul_eq_zero_iff_left (hx : x ≠ 0) : c • x = 0 ↔ c = 0 := by simp [hx] +lemma smul_eq_zero_iff_right (hc : c ≠ 0) : c • x = 0 ↔ x = 0 := by simp [hc] +lemma smul_ne_zero_iff_left (hx : x ≠ 0) : c • x ≠ 0 ↔ c ≠ 0 := by simp [hx] +lemma smul_ne_zero_iff_right (hc : c ≠ 0) : c • x ≠ 0 ↔ x ≠ 0 := by simp [hc] + +end SMulWithZero + +end NoZeroSMulDivisors diff --git a/Mathlib/Algebra/NoZeroSMulDivisors/Pi.lean b/Mathlib/Algebra/NoZeroSMulDivisors/Pi.lean new file mode 100644 index 0000000000000..ac7b42a019040 --- /dev/null +++ b/Mathlib/Algebra/NoZeroSMulDivisors/Pi.lean @@ -0,0 +1,34 @@ +/- +Copyright (c) 2018 Simon Hudon. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Anne Baanen, Yaël Dillies +-/ +import Mathlib.Algebra.NoZeroSMulDivisors.Defs +import Mathlib.Algebra.Group.Action.Pi + +/-! +# Pi instances for NoZeroSMulDivisors + +This file defines instances for NoZeroSMulDivisors on Pi types. +-/ + + +universe u v + +variable {I : Type u} + +-- The indexing type +variable {f : I → Type v} + +instance Pi.noZeroSMulDivisors (α) [Zero α] [∀ i, Zero <| f i] + [∀ i, SMulWithZero α <| f i] [∀ i, NoZeroSMulDivisors α <| f i] : + NoZeroSMulDivisors α (∀ i : I, f i) := + ⟨fun {_ _} h => + or_iff_not_imp_left.mpr fun hc => + funext fun i => (smul_eq_zero.mp (congr_fun h i)).resolve_left hc⟩ + +/-- A special case of `Pi.noZeroSMulDivisors` for non-dependent types. Lean struggles to +synthesize this instance by itself elsewhere in the library. -/ +instance _root_.Function.noZeroSMulDivisors {ι α β : Type*} [Zero α] [Zero β] + [SMulWithZero α β] [NoZeroSMulDivisors α β] : NoZeroSMulDivisors α (ι → β) := + Pi.noZeroSMulDivisors _ diff --git a/Mathlib/Algebra/NoZeroSMulDivisors/Prod.lean b/Mathlib/Algebra/NoZeroSMulDivisors/Prod.lean new file mode 100644 index 0000000000000..e836979c0377a --- /dev/null +++ b/Mathlib/Algebra/NoZeroSMulDivisors/Prod.lean @@ -0,0 +1,32 @@ +/- +Copyright (c) 2018 Simon Hudon. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Anne Baanen +-/ +import Mathlib.Algebra.NoZeroSMulDivisors.Defs +import Mathlib.Algebra.Group.Action.Prod + +/-! +# Prod instances for NoZeroSMulDivisors + +This file defines a NoZeroSMulDivisors instance for the binary product of actions. +-/ + +variable {R M N : Type*} + +namespace Prod + +instance noZeroSMulDivisors [Zero R] [Zero M] [Zero N] + [SMulWithZero R M] [SMulWithZero R N] [NoZeroSMulDivisors R M] [NoZeroSMulDivisors R N] : + NoZeroSMulDivisors R (M × N) := + { eq_zero_or_eq_zero_of_smul_eq_zero := by -- Porting note: in mathlib3 there is no need for `by`/ + -- `intro`/`exact`, i.e. the following works: + -- ⟨fun c ⟨x, y⟩ h => + -- or_iff_not_imp_left.mpr fun hc => + intro c ⟨x, y⟩ h + exact or_iff_not_imp_left.mpr fun hc => + mk.inj_iff.mpr + ⟨(smul_eq_zero.mp (congr_arg fst h)).resolve_left hc, + (smul_eq_zero.mp (congr_arg snd h)).resolve_left hc⟩ } + +end Prod diff --git a/Mathlib/Algebra/Order/Module/Defs.lean b/Mathlib/Algebra/Order/Module/Defs.lean index 8323669b29434..a1172e5a8524d 100644 --- a/Mathlib/Algebra/Order/Module/Defs.lean +++ b/Mathlib/Algebra/Order/Module/Defs.lean @@ -3,7 +3,7 @@ Copyright (c) 2023 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ -import Mathlib.Algebra.Module.Defs +import Mathlib.Algebra.NoZeroSMulDivisors.Basic import Mathlib.Algebra.Order.Field.Defs import Mathlib.Algebra.Order.GroupWithZero.Action.Synonym import Mathlib.Tactic.GCongr diff --git a/Mathlib/Algebra/Order/Star/Basic.lean b/Mathlib/Algebra/Order/Star/Basic.lean index 6370c2c91944e..6df890e8638a3 100644 --- a/Mathlib/Algebra/Order/Star/Basic.lean +++ b/Mathlib/Algebra/Order/Star/Basic.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Kim Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison -/ +import Mathlib.Algebra.NoZeroSMulDivisors.Defs import Mathlib.Algebra.Order.Group.Defs import Mathlib.Algebra.Order.Group.Nat import Mathlib.Algebra.Star.SelfAdjoint diff --git a/Mathlib/Data/ENNReal/Operations.lean b/Mathlib/Data/ENNReal/Operations.lean index 473c3d6c65d78..26c974025ccf6 100644 --- a/Mathlib/Data/ENNReal/Operations.lean +++ b/Mathlib/Data/ENNReal/Operations.lean @@ -5,7 +5,6 @@ Authors: Johannes Hölzl, Yury Kudryashov -/ import Mathlib.Algebra.BigOperators.WithTop import Mathlib.Algebra.GroupWithZero.Divisibility -import Mathlib.Algebra.Module.Basic import Mathlib.Data.ENNReal.Basic /-! diff --git a/Mathlib/Data/Set/Pointwise/SMul.lean b/Mathlib/Data/Set/Pointwise/SMul.lean index 9e6917fe0172c..bd562b3fb02ad 100644 --- a/Mathlib/Data/Set/Pointwise/SMul.lean +++ b/Mathlib/Data/Set/Pointwise/SMul.lean @@ -7,6 +7,7 @@ import Mathlib.Algebra.Group.Pi.Basic import Mathlib.Algebra.Group.Pointwise.Set.Basic import Mathlib.Algebra.GroupWithZero.Action.Basic import Mathlib.Algebra.Module.Defs +import Mathlib.Algebra.NoZeroSMulDivisors.Defs import Mathlib.Data.Set.Pairwise.Basic /-! diff --git a/Mathlib/GroupTheory/Subgroup/Saturated.lean b/Mathlib/GroupTheory/Subgroup/Saturated.lean index 1d7d366b914ad..6cbc7b2c45d71 100644 --- a/Mathlib/GroupTheory/Subgroup/Saturated.lean +++ b/Mathlib/GroupTheory/Subgroup/Saturated.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin -/ import Mathlib.Algebra.Group.Subgroup.Basic -import Mathlib.Algebra.Module.Defs +import Mathlib.Algebra.NoZeroSMulDivisors.Defs /-! # Saturated subgroups diff --git a/Mathlib/LinearAlgebra/BilinearMap.lean b/Mathlib/LinearAlgebra/BilinearMap.lean index 0430d8a20645b..c720aff5ff9b6 100644 --- a/Mathlib/LinearAlgebra/BilinearMap.lean +++ b/Mathlib/LinearAlgebra/BilinearMap.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau, Mario Carneiro -/ import Mathlib.Algebra.Module.Submodule.Ker +import Mathlib.Algebra.NoZeroSMulDivisors.Basic /-! # Basics on bilinear maps diff --git a/Mathlib/LinearAlgebra/LinearPMap.lean b/Mathlib/LinearAlgebra/LinearPMap.lean index ff0ddc27cc274..6cc05feeabb39 100644 --- a/Mathlib/LinearAlgebra/LinearPMap.lean +++ b/Mathlib/LinearAlgebra/LinearPMap.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov, Moritz Doll -/ import Mathlib.LinearAlgebra.Prod -import Mathlib.Algebra.Module.Basic /-! # Partially defined linear maps diff --git a/Mathlib/LinearAlgebra/Multilinear/Basic.lean b/Mathlib/LinearAlgebra/Multilinear/Basic.lean index 8db7f46936b45..ea1242276149d 100644 --- a/Mathlib/LinearAlgebra/Multilinear/Basic.lean +++ b/Mathlib/LinearAlgebra/Multilinear/Basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ import Mathlib.Algebra.Algebra.Defs +import Mathlib.Algebra.NoZeroSMulDivisors.Pi import Mathlib.Algebra.Order.BigOperators.Group.Finset import Mathlib.Data.Fintype.BigOperators import Mathlib.Data.Fintype.Sort diff --git a/Mathlib/LinearAlgebra/Span.lean b/Mathlib/LinearAlgebra/Span.lean index 8e873cfc1c533..478735c139125 100644 --- a/Mathlib/LinearAlgebra/Span.lean +++ b/Mathlib/LinearAlgebra/Span.lean @@ -8,6 +8,7 @@ import Mathlib.Algebra.Module.Prod import Mathlib.Algebra.Module.Submodule.EqLocus import Mathlib.Algebra.Module.Submodule.Equiv import Mathlib.Algebra.Module.Submodule.RestrictScalars +import Mathlib.Algebra.NoZeroSMulDivisors.Basic import Mathlib.Algebra.Ring.Idempotents import Mathlib.Data.Set.Pointwise.SMul import Mathlib.Order.CompactlyGenerated.Basic diff --git a/Mathlib/RingTheory/Nilpotent/Basic.lean b/Mathlib/RingTheory/Nilpotent/Basic.lean index 71e196c394db9..9eddb6417010a 100644 --- a/Mathlib/RingTheory/Nilpotent/Basic.lean +++ b/Mathlib/RingTheory/Nilpotent/Basic.lean @@ -7,7 +7,7 @@ import Mathlib.Algebra.Associated.Basic import Mathlib.Algebra.GeomSum import Mathlib.Algebra.Group.Action.Prod import Mathlib.Algebra.GroupWithZero.NonZeroDivisors -import Mathlib.Algebra.Module.Defs +import Mathlib.Algebra.NoZeroSMulDivisors.Defs import Mathlib.Algebra.SMulWithZero import Mathlib.Data.Nat.Choose.Sum import Mathlib.Data.Nat.Lattice From 418a5eb7aec3fb639097cb13f74fc031ac4057f2 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Sat, 19 Oct 2024 07:30:34 +0000 Subject: [PATCH 178/505] chore: remove CoeFun instances where FunLike is available (#17911) During the port we found that `FunLike` is robust enough not to need an extra `CoeFun` shortcut. Let's see if that rule can be consistently applied to the whole of the library. There is still duplication between `FunLike` and `CoeFun` for `Grp`, `Mon`, `CommGrp` and `CommMon`, which will need a more thorough fix. See also #17866. I am currently bisecting to figure out exactly why the benchmarks are indicating a slowdown with no obvious cause. --- Mathlib/Algebra/Order/AbsoluteValue.lean | 5 ----- Mathlib/Algebra/Ring/CentroidHom.lean | 7 ------- .../Analysis/Distribution/SchwartzSpace.lean | 4 ---- Mathlib/Analysis/Normed/Algebra/Norm.lean | 8 ------- Mathlib/Analysis/Normed/Group/Seminorm.lean | 21 ------------------- 5 files changed, 45 deletions(-) diff --git a/Mathlib/Algebra/Order/AbsoluteValue.lean b/Mathlib/Algebra/Order/AbsoluteValue.lean index 0dcc6acdcfbb4..fae16b0ffcd6d 100644 --- a/Mathlib/Algebra/Order/AbsoluteValue.lean +++ b/Mathlib/Algebra/Order/AbsoluteValue.lean @@ -77,11 +77,6 @@ def Simps.apply (f : AbsoluteValue R S) : R → S := initialize_simps_projections AbsoluteValue (toMulHom_toFun → apply) -/-- Helper instance for when there's too many metavariables to apply `DFunLike.has_coe_to_fun` -directly. -/ -instance : CoeFun (AbsoluteValue R S) fun _ => R → S := - DFunLike.hasCoeToFun - @[simp] theorem coe_toMulHom : ⇑abv.toMulHom = abv := rfl diff --git a/Mathlib/Algebra/Ring/CentroidHom.lean b/Mathlib/Algebra/Ring/CentroidHom.lean index 37a70431efb16..b1b90bb34b9d1 100644 --- a/Mathlib/Algebra/Ring/CentroidHom.lean +++ b/Mathlib/Algebra/Ring/CentroidHom.lean @@ -102,13 +102,6 @@ instance : CentroidHomClass (CentroidHom α) α where map_mul_right f := f.map_mul_right' -/-- Helper instance for when there's too many metavariables to apply `DFunLike.CoeFun` -directly. -/ -/- Porting note: Lean gave me `unknown constant 'DFunLike.CoeFun'` and says `CoeFun` is a type -mismatch, so I used `library_search`. -/ -instance : CoeFun (CentroidHom α) fun _ ↦ α → α := - inferInstanceAs (CoeFun (CentroidHom α) fun _ ↦ α → α) - -- Porting note: removed @[simp]; not in normal form. (`toAddMonoidHom_eq_coe` below ensures that -- the LHS simplifies to the RHS anyway.) theorem toFun_eq_coe {f : CentroidHom α} : f.toFun = f := rfl diff --git a/Mathlib/Analysis/Distribution/SchwartzSpace.lean b/Mathlib/Analysis/Distribution/SchwartzSpace.lean index ba8aae073aff9..a9cd6e02b3fbe 100644 --- a/Mathlib/Analysis/Distribution/SchwartzSpace.lean +++ b/Mathlib/Analysis/Distribution/SchwartzSpace.lean @@ -90,10 +90,6 @@ instance instFunLike : FunLike 𝓢(E, F) E F where coe f := f.toFun coe_injective' f g h := by cases f; cases g; congr -/-- Helper instance for when there's too many metavariables to apply `DFunLike.hasCoeToFun`. -/ -instance instCoeFun : CoeFun 𝓢(E, F) fun _ => E → F := - DFunLike.hasCoeToFun - /-- All derivatives of a Schwartz function are rapidly decaying. -/ theorem decay (f : 𝓢(E, F)) (k n : ℕ) : ∃ C : ℝ, 0 < C ∧ ∀ x, ‖x‖ ^ k * ‖iteratedFDeriv ℝ n f x‖ ≤ C := by diff --git a/Mathlib/Analysis/Normed/Algebra/Norm.lean b/Mathlib/Analysis/Normed/Algebra/Norm.lean index eb815ac91298e..715ad8694cf6f 100644 --- a/Mathlib/Analysis/Normed/Algebra/Norm.lean +++ b/Mathlib/Analysis/Normed/Algebra/Norm.lean @@ -70,10 +70,6 @@ instance algebraNormClass : AlgebraNormClass (AlgebraNorm R S) R S where eq_zero_of_map_eq_zero f := f.eq_zero_of_map_eq_zero' _ map_smul_eq_mul f := f.smul' -/-- Helper instance for when there's too many metavariables to apply `fun_like.has_coe_to_fun`. -/ -instance : CoeFun (AlgebraNorm R S) fun _ => S → ℝ := - DFunLike.hasCoeToFun - theorem toFun_eq_coe (p : AlgebraNorm R S) : p.toFun = p := rfl @[ext] @@ -160,10 +156,6 @@ instance mulAlgebraNormClass : MulAlgebraNormClass (MulAlgebraNorm R S) R S wher eq_zero_of_map_eq_zero f := f.eq_zero_of_map_eq_zero' _ map_smul_eq_mul f := f.smul' -/-- Helper instance for when there's too many metavariables to apply `fun_like.has_coe_to_fun`. -/ -instance : CoeFun (MulAlgebraNorm R S) fun _ => S → ℝ := - DFunLike.hasCoeToFun - theorem toFun_eq_coe (p : MulAlgebraNorm R S) : p.toFun = p := rfl @[ext] diff --git a/Mathlib/Analysis/Normed/Group/Seminorm.lean b/Mathlib/Analysis/Normed/Group/Seminorm.lean index eede911fb046d..a8d9220ba89cc 100644 --- a/Mathlib/Analysis/Normed/Group/Seminorm.lean +++ b/Mathlib/Analysis/Normed/Group/Seminorm.lean @@ -181,12 +181,6 @@ instance groupSeminormClass : GroupSeminormClass (GroupSeminorm E) E ℝ where map_mul_le_add f := f.mul_le' map_inv_eq_map f := f.inv' -/-- Helper instance for when there's too many metavariables to apply `DFunLike.hasCoeToFun`. -/ -@[to_additive "Helper instance for when there's too many metavariables to apply -`DFunLike.hasCoeToFun`. "] -instance : CoeFun (GroupSeminorm E) fun _ => E → ℝ := - ⟨DFunLike.coe⟩ - @[to_additive (attr := simp)] theorem toFun_eq_coe : p.toFun = p := rfl @@ -447,10 +441,6 @@ instance nonarchAddGroupSeminormClass : map_zero f := f.map_zero' map_neg_eq_map' f := f.neg' -/-- Helper instance for when there's too many metavariables to apply `DFunLike.hasCoeToFun`. -/ -instance : CoeFun (NonarchAddGroupSeminorm E) fun _ => E → ℝ := - ⟨DFunLike.coe⟩ - -- Porting note: `simpNF` said the left hand side simplified to this @[simp] theorem toZeroHom_eq_coe : ⇑p.toZeroHom = p := by @@ -667,13 +657,6 @@ instance groupNormClass : GroupNormClass (GroupNorm E) E ℝ where map_inv_eq_map f := f.inv' eq_one_of_map_eq_zero f := f.eq_one_of_map_eq_zero' _ -/-- Helper instance for when there's too many metavariables to apply `DFunLike.hasCoeToFun` -directly. -/ -@[to_additive "Helper instance for when there's too many metavariables to apply -`DFunLike.hasCoeToFun` directly. "] -instance : CoeFun (GroupNorm E) fun _ => E → ℝ := - DFunLike.hasCoeToFun - -- Porting note: `simpNF` told me the left-hand side simplified to this @[to_additive (attr := simp)] theorem toGroupSeminorm_eq_coe : ⇑p.toGroupSeminorm = p := @@ -799,10 +782,6 @@ instance nonarchAddGroupNormClass : NonarchAddGroupNormClass (NonarchAddGroupNor map_neg_eq_map' f := f.neg' eq_zero_of_map_eq_zero f := f.eq_zero_of_map_eq_zero' _ -/-- Helper instance for when there's too many metavariables to apply `DFunLike.hasCoeToFun`. -/ -noncomputable instance : CoeFun (NonarchAddGroupNorm E) fun _ => E → ℝ := - DFunLike.hasCoeToFun - -- Porting note: `simpNF` told me the left-hand side simplified to this @[simp] theorem toNonarchAddGroupSeminorm_eq_coe : ⇑p.toNonarchAddGroupSeminorm = p := From 66f63643834286e06540e60e459b22c7bbd25475 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Sat, 19 Oct 2024 08:15:11 +0000 Subject: [PATCH 179/505] chore: move trans tactic to Batteries (batteries#1001) (#17931) Co-authored-by: F. G. Dorais --- Mathlib.lean | 1 - Mathlib/Algebra/Notation.lean | 1 + Mathlib/Order/Defs.lean | 2 +- Mathlib/Tactic.lean | 1 - Mathlib/Tactic/Common.lean | 1 - Mathlib/Tactic/MoveAdd.lean | 1 + Mathlib/Tactic/Relation/Trans.lean | 222 ------------------------ Mathlib/Tactic/ToAdditive/Frontend.lean | 4 +- Mathlib/Tactic/Widget/CongrM.lean | 1 + lake-manifest.json | 2 +- scripts/noshake.json | 2 +- test/trans.lean | 109 ------------ 12 files changed, 8 insertions(+), 339 deletions(-) delete mode 100644 Mathlib/Tactic/Relation/Trans.lean delete mode 100644 test/trans.lean diff --git a/Mathlib.lean b/Mathlib.lean index 66e48c38b8ae0..e24762ba7699b 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4447,7 +4447,6 @@ import Mathlib.Tactic.ReduceModChar import Mathlib.Tactic.ReduceModChar.Ext import Mathlib.Tactic.Relation.Rfl import Mathlib.Tactic.Relation.Symm -import Mathlib.Tactic.Relation.Trans import Mathlib.Tactic.Rename import Mathlib.Tactic.RenameBVar import Mathlib.Tactic.Replace diff --git a/Mathlib/Algebra/Notation.lean b/Mathlib/Algebra/Notation.lean index 0ea472b668eb9..ffb3097c7ba73 100644 --- a/Mathlib/Algebra/Notation.lean +++ b/Mathlib/Algebra/Notation.lean @@ -3,6 +3,7 @@ Copyright (c) 2024 Jireh Loreaux. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jireh Loreaux -/ +import Mathlib.Tactic.TypeStar import Mathlib.Tactic.ToAdditive /-! diff --git a/Mathlib/Order/Defs.lean b/Mathlib/Order/Defs.lean index 7a46f0ce76cf7..75d603ee2c080 100644 --- a/Mathlib/Order/Defs.lean +++ b/Mathlib/Order/Defs.lean @@ -4,9 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura -/ import Batteries.Classes.Order +import Batteries.Tactic.Trans import Mathlib.Data.Ordering.Basic import Mathlib.Tactic.Lemma -import Mathlib.Tactic.Relation.Trans import Mathlib.Tactic.SplitIfs import Mathlib.Tactic.TypeStar diff --git a/Mathlib/Tactic.lean b/Mathlib/Tactic.lean index cd494193a67b4..efac05b1b7da1 100644 --- a/Mathlib/Tactic.lean +++ b/Mathlib/Tactic.lean @@ -203,7 +203,6 @@ import Mathlib.Tactic.ReduceModChar import Mathlib.Tactic.ReduceModChar.Ext import Mathlib.Tactic.Relation.Rfl import Mathlib.Tactic.Relation.Symm -import Mathlib.Tactic.Relation.Trans import Mathlib.Tactic.Rename import Mathlib.Tactic.RenameBVar import Mathlib.Tactic.Replace diff --git a/Mathlib/Tactic/Common.lean b/Mathlib/Tactic/Common.lean index f9c7a25eb782c..c864e4fc60f26 100644 --- a/Mathlib/Tactic/Common.lean +++ b/Mathlib/Tactic/Common.lean @@ -79,7 +79,6 @@ import Mathlib.Tactic.PushNeg import Mathlib.Tactic.RSuffices import Mathlib.Tactic.Recover import Mathlib.Tactic.Relation.Rfl -import Mathlib.Tactic.Relation.Trans import Mathlib.Tactic.Rename import Mathlib.Tactic.RenameBVar import Mathlib.Tactic.Says diff --git a/Mathlib/Tactic/MoveAdd.lean b/Mathlib/Tactic/MoveAdd.lean index 094ce46b60265..87cf3bdb6add4 100644 --- a/Mathlib/Tactic/MoveAdd.lean +++ b/Mathlib/Tactic/MoveAdd.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Arthur Paulino, Damiano Testa -/ import Mathlib.Algebra.Group.Basic +import Mathlib.Lean.Meta /-! diff --git a/Mathlib/Tactic/Relation/Trans.lean b/Mathlib/Tactic/Relation/Trans.lean deleted file mode 100644 index 290d23e9aeb70..0000000000000 --- a/Mathlib/Tactic/Relation/Trans.lean +++ /dev/null @@ -1,222 +0,0 @@ -/- -Copyright (c) 2022 Siddhartha Gadgil. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Siddhartha Gadgil, Mario Carneiro --/ -import Mathlib.Lean.Meta -import Mathlib.Lean.Elab.Tactic.Basic -import Lean.Elab.Tactic.ElabTerm -import Mathlib.Tactic.TypeStar - -/-! -# `trans` tactic - -This implements the `trans` tactic, which can apply transitivity theorems with an optional middle -variable argument. --/ - -namespace Mathlib.Tactic -open Lean Meta Elab - -initialize registerTraceClass `Tactic.trans - -/-- Discrimation tree settings for the `trans` extension. -/ -def transExt.config : WhnfCoreConfig := {} - -/-- Environment extension storing transitivity lemmas -/ -initialize transExt : - SimpleScopedEnvExtension (Name × Array DiscrTree.Key) (DiscrTree Name) ← - registerSimpleScopedEnvExtension { - addEntry := fun dt (n, ks) ↦ dt.insertCore ks n - initial := {} - } - -initialize registerBuiltinAttribute { - name := `trans - descr := "transitive relation" - add := fun decl _ kind ↦ MetaM.run' do - let declTy := (← getConstInfo decl).type - let (xs, _, targetTy) ← withReducible <| forallMetaTelescopeReducing declTy - let fail := throwError - "@[trans] attribute only applies to lemmas proving - x ∼ y → y ∼ z → x ∼ z, got {indentExpr declTy} with target {indentExpr targetTy}" - let .app (.app rel _) _ := targetTy | fail - let some yzHyp := xs.back? | fail - let some xyHyp := xs.pop.back? | fail - let .app (.app _ _) _ ← inferType yzHyp | fail - let .app (.app _ _) _ ← inferType xyHyp | fail - let key ← withReducible <| DiscrTree.mkPath rel transExt.config - transExt.add (decl, key) kind -} - -universe u v in -/-- Composition using the `Trans` class in the homogeneous case. -/ -def _root_.Trans.simple {α : Sort u} {r : α → α → Sort v} {a b c : α} [Trans r r r] : - r a b → r b c → r a c := trans - -universe u v w in -/-- Composition using the `Trans` class in the general case. -/ -def _root_.Trans.het {α β γ : Sort*} {a : α} {b : β} {c : γ} - {r : α → β → Sort u} {s : β → γ → Sort v} {t : outParam (α → γ → Sort w)} [Trans r s t] : - r a b → s b c → t a c := trans - -open Lean.Elab.Tactic - -/-- solving `e ← mkAppM' f #[x]` -/ -def getExplicitFuncArg? (e : Expr) : MetaM (Option <| Expr × Expr) := do - match e with - | Expr.app f a => do - if ← isDefEq (← mkAppM' f #[a]) e then - return some (f, a) - else - getExplicitFuncArg? f - | _ => return none - -/-- solving `tgt ← mkAppM' rel #[x, z]` given `tgt = f z` -/ -def getExplicitRelArg? (tgt f z : Expr) : MetaM (Option <| Expr × Expr) := do - match f with - | Expr.app rel x => do - let check: Bool ← do - try - let folded ← mkAppM' rel #[x, z] - isDefEq folded tgt - catch _ => - pure false - if check then - return some (rel, x) - else - getExplicitRelArg? tgt rel z - | _ => return none - -/-- refining `tgt ← mkAppM' rel #[x, z]` dropping more arguments if possible -/ -def getExplicitRelArgCore (tgt rel x z : Expr) : MetaM (Expr × Expr) := do - match rel with - | Expr.app rel' _ => do - let check: Bool ← do - try - let folded ← mkAppM' rel' #[x, z] - isDefEq folded tgt - catch _ => - pure false - if !check then - return (rel, x) - else - getExplicitRelArgCore tgt rel' x z - | _ => return (rel ,x) - -/-- Internal definition for `trans` tactic. Either a binary relation or a non-dependent -arrow. -/ -inductive TransRelation - | app (rel : Expr) - | implies (name : Name) (bi : BinderInfo) - -/-- Finds an explicit binary relation in the argument, if possible. -/ -def getRel (tgt : Expr) : MetaM (Option (TransRelation × Expr × Expr)) := do - match tgt with - | .forallE name binderType body info => return .some (.implies name info, binderType, body) - | .app f z => - match (← getExplicitRelArg? tgt f z) with - | some (rel, x) => - let (rel, x) ← getExplicitRelArgCore tgt rel x z - return some (.app rel, x, z) - | none => - return none - | _ => return none - -/-- -`trans` applies to a goal whose target has the form `t ~ u` where `~` is a transitive relation, -that is, a relation which has a transitivity lemma tagged with the attribute [trans]. - -* `trans s` replaces the goal with the two subgoals `t ~ s` and `s ~ u`. -* If `s` is omitted, then a metavariable is used instead. - -Additionally, `trans` also applies to a goal whose target has the form `t → u`, -in which case it replaces the goal with `t → s` and `s → u`. --/ -elab "trans" t?:(ppSpace colGt term)? : tactic => withMainContext do - let tgt ← getMainTarget'' - let .some (rel, x, z) ← getRel tgt | - throwError (m!"transitivity lemmas only apply to binary relations and " ++ - m!"non-dependent arrows, not {indentExpr tgt}") - match rel with - | .implies name info => - -- only consider non-dependent arrows - if z.hasLooseBVars then - throwError "`trans` is not implemented for dependent arrows{indentExpr tgt}" - -- parse the intermeditate term - let middleType ← mkFreshExprMVar none - let t'? ← t?.mapM (elabTermWithHoles · middleType (← getMainTag)) - let middle ← (t'?.map (pure ·.1)).getD (mkFreshExprMVar middleType) - liftMetaTactic fun goal => do - -- create two new goals - let g₁ ← mkFreshExprMVar (some <| .forallE name x middle info) .synthetic - let g₂ ← mkFreshExprMVar (some <| .forallE name middle z info) .synthetic - -- close the original goal with `fun x => g₂ (g₁ x)` - goal.assign (.lam name x (.app g₂ (.app g₁ (.bvar 0))) .default) - pure <| [g₁.mvarId!, g₂.mvarId!] ++ if let some (_, gs') := t'? then gs' else [middle.mvarId!] - return - | .app rel => - trace[Tactic.trans]"goal decomposed" - trace[Tactic.trans]"rel: {indentExpr rel}" - trace[Tactic.trans]"x: {indentExpr x}" - trace[Tactic.trans]"z: {indentExpr z}" - -- first trying the homogeneous case - try - let ty ← inferType x - let t'? ← t?.mapM (elabTermWithHoles · ty (← getMainTag)) - let s ← saveState - trace[Tactic.trans]"trying homogeneous case" - let lemmas := - (← (transExt.getState (← getEnv)).getUnify rel transExt.config).push ``Trans.simple - for lem in lemmas do - trace[Tactic.trans]"trying lemma {lem}" - try - liftMetaTactic fun g ↦ do - let lemTy ← inferType (← mkConstWithLevelParams lem) - let arity ← withReducible <| forallTelescopeReducing lemTy fun es _ ↦ pure es.size - let y ← (t'?.map (pure ·.1)).getD (mkFreshExprMVar ty) - let g₁ ← mkFreshExprMVar (some <| ← mkAppM' rel #[x, y]) .synthetic - let g₂ ← mkFreshExprMVar (some <| ← mkAppM' rel #[y, z]) .synthetic - g.assign (← mkAppOptM lem (mkArray (arity - 2) none ++ #[some g₁, some g₂])) - pure <| [g₁.mvarId!, g₂.mvarId!] ++ - if let some (_, gs') := t'? then gs' else [y.mvarId!] - return - catch _ => s.restore - pure () - catch _ => - trace[Tactic.trans]"trying heterogeneous case" - let t'? ← t?.mapM (elabTermWithHoles · none (← getMainTag)) - let s ← saveState - for lem in (← (transExt.getState (← getEnv)).getUnify rel transExt.config).push - ``HEq.trans |>.push ``HEq.trans do - try - liftMetaTactic fun g ↦ do - trace[Tactic.trans]"trying lemma {lem}" - let lemTy ← inferType (← mkConstWithLevelParams lem) - let arity ← withReducible <| forallTelescopeReducing lemTy fun es _ ↦ pure es.size - trace[Tactic.trans]"arity: {arity}" - trace[Tactic.trans]"lemma-type: {lemTy}" - let y ← (t'?.map (pure ·.1)).getD (mkFreshExprMVar none) - trace[Tactic.trans]"obtained y: {y}" - trace[Tactic.trans]"rel: {indentExpr rel}" - trace[Tactic.trans]"x:{indentExpr x}" - trace[Tactic.trans]"z: {indentExpr z}" - let g₂ ← mkFreshExprMVar (some <| ← mkAppM' rel #[y, z]) .synthetic - trace[Tactic.trans]"obtained g₂: {g₂}" - let g₁ ← mkFreshExprMVar (some <| ← mkAppM' rel #[x, y]) .synthetic - trace[Tactic.trans]"obtained g₁: {g₁}" - g.assign (← mkAppOptM lem (mkArray (arity - 2) none ++ #[some g₁, some g₂])) - pure <| [g₁.mvarId!, g₂.mvarId!] ++ if let some (_, gs') := t'? then gs' else [y.mvarId!] - return - catch e => - trace[Tactic.trans]"failed: {e.toMessageData}" - s.restore - throwError m!"no applicable transitivity lemma found for {indentExpr tgt}" - -syntax "transitivity" (ppSpace colGt term)? : tactic -set_option hygiene false in -macro_rules - | `(tactic| transitivity) => `(tactic| trans) - | `(tactic| transitivity $e) => `(tactic| trans $e) - -end Mathlib.Tactic diff --git a/Mathlib/Tactic/ToAdditive/Frontend.lean b/Mathlib/Tactic/ToAdditive/Frontend.lean index 277e01a289f6b..c0d0ef7afdcfa 100644 --- a/Mathlib/Tactic/ToAdditive/Frontend.lean +++ b/Mathlib/Tactic/ToAdditive/Frontend.lean @@ -16,7 +16,7 @@ import Lean.Meta.Tactic.Rfl import Lean.Meta.Match.MatcherInfo import Batteries.Lean.NameMapAttribute import Batteries.Tactic.Lint -- useful to lint this file and for DiscrTree.elements -import Mathlib.Tactic.Relation.Trans -- just to copy the attribute +import Batteries.Tactic.Trans import Mathlib.Tactic.Eqns -- just to copy the attribute import Mathlib.Tactic.Simps.Basic @@ -1167,7 +1167,7 @@ partial def applyAttributes (stx : Syntax) (rawAttrs : Array Syntax) (thisAttr s (fun b n => (b.tree.values.any fun t => t.declName = n)) thisAttr `ext src tgt warnAttr stx Lean.Meta.Rfl.reflExt (·.values.contains ·) thisAttr `refl src tgt warnAttr stx Lean.Meta.Symm.symmExt (·.values.contains ·) thisAttr `symm src tgt - warnAttr stx Mathlib.Tactic.transExt (·.values.contains ·) thisAttr `trans src tgt + warnAttr stx Batteries.Tactic.transExt (·.values.contains ·) thisAttr `trans src tgt warnAttr stx Lean.Meta.coeExt (·.contains ·) thisAttr `coe src tgt warnParametricAttr stx Lean.Linter.deprecatedAttr thisAttr `deprecated src tgt -- the next line also warns for `@[to_additive, simps]`, because of the application times diff --git a/Mathlib/Tactic/Widget/CongrM.lean b/Mathlib/Tactic/Widget/CongrM.lean index 9c259c3daaf30..dad4520d652cc 100644 --- a/Mathlib/Tactic/Widget/CongrM.lean +++ b/Mathlib/Tactic/Widget/CongrM.lean @@ -6,6 +6,7 @@ Authors: Patrick Massot import Mathlib.Tactic.Widget.SelectPanelUtils import Mathlib.Tactic.CongrM +import Batteries.Lean.Position /-! # CongrM widget diff --git a/lake-manifest.json b/lake-manifest.json index 6f116ba1524dd..8d8fe8356466a 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,7 +5,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "1e0bf50b357069e1d658512a579a5faac6587c40", + "rev": "10130798199d306703dee5ab2567961444ebbd04", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", diff --git a/scripts/noshake.json b/scripts/noshake.json index 3e93d378ca0e2..a98fca062e875 100644 --- a/scripts/noshake.json +++ b/scripts/noshake.json @@ -148,7 +148,6 @@ "Mathlib.Tactic.Recover", "Mathlib.Tactic.ReduceModChar.Ext", "Mathlib.Tactic.Relation.Symm", - "Mathlib.Tactic.Relation.Trans", "Mathlib.Tactic.Rename", "Mathlib.Tactic.RenameBVar", "Mathlib.Tactic.Replace", @@ -325,6 +324,7 @@ "Mathlib.Order.RelClasses": ["Batteries.WF"], "Mathlib.Order.Filter.ListTraverse": ["Mathlib.Control.Traversable.Instances"], + "Mathlib.Order.Defs": ["Batteries.Tactic.Trans"], "Mathlib.Order.Basic": ["Batteries.Tactic.Classical"], "Mathlib.NumberTheory.Harmonic.Defs": ["Mathlib.Algebra.Order.BigOperators.Ring.Finset", diff --git a/test/trans.lean b/test/trans.lean deleted file mode 100644 index 50b5bc01e7b20..0000000000000 --- a/test/trans.lean +++ /dev/null @@ -1,109 +0,0 @@ -import Mathlib.Tactic.Relation.Trans -import Mathlib.Tactic.GuardGoalNums - -set_option autoImplicit true --- testing that the attribute is recognized and used -def nleq (a b : Nat) : Prop := a ≤ b - -@[trans] def nleq_trans : nleq a b → nleq b c → nleq a c := Nat.le_trans - -example (a b c : Nat) : nleq a b → nleq b c → nleq a c := by - intro h₁ h₂ - trans b - assumption - assumption - -example (a b c : Nat) : nleq a b → nleq b c → nleq a c := by intros; trans <;> assumption - --- using `Trans` typeclass -@[trans] def eq_trans {a b c : α} : a = b → b = c → a = c := by - intro h₁ h₂ - apply Eq.trans h₁ h₂ - -example (a b c : Nat) : a = b → b = c → a = c := by intros; trans <;> assumption - -example (a b c : Nat) : a = b → b = c → a = c := by - intro h₁ h₂ - trans b - assumption - assumption - -example : @Trans Nat Nat Nat (· ≤ ·) (· ≤ ·) (· ≤ ·) := inferInstance - -example (a b c : Nat) : a ≤ b → b ≤ c → a ≤ c := by - intros h₁ h₂ - trans ?b - case b => exact b - exact h₁ - exact h₂ - -example (a b c : α) (R : α → α → Prop) [Trans R R R] : R a b → R b c → R a c := by - intros h₁ h₂ - trans ?b - case b => exact b - exact h₁ - exact h₂ - -example (a b c : Nat) : a ≤ b → b ≤ c → a ≤ c := by - intros h₁ h₂ - trans - exact h₁ - exact h₂ - -example (a b c : Nat) : a ≤ b → b ≤ c → a ≤ c := by intros; trans <;> assumption - -example (a b c : Nat) : a < b → b < c → a < c := by - intro h₁ h₂ - trans b - assumption - assumption - -example (a b c : Nat) : a < b → b < c → a < c := by intros; trans <;> assumption - -example (x n p : Nat) (h₁ : n * Nat.succ p ≤ x) : n * p ≤ x := by - trans - · apply Nat.mul_le_mul_left; apply Nat.le_succ - · apply h₁ - -example (a : α) (c : γ) : ∀ b : β, HEq a b → HEq b c → HEq a c := by - intro b h₁ h₂ - trans b - assumption - assumption - -def MyLE (n m : Nat) := ∃ k, n + k = m - -@[trans] theorem MyLE.trans {n m k : Nat} (h1 : MyLE n m) (h2 : MyLE m k) : MyLE n k := by - cases h1 - cases h2 - subst_vars - exact ⟨_, Eq.symm <| Nat.add_assoc _ _ _⟩ - -example {n m k : Nat} (h1 : MyLE n m) (h2 : MyLE m k) : MyLE n k := by - trans <;> assumption - -/-- `trans` for implications. -/ -example {A B C : Prop} (h : A → B) (g : B → C) : A → C := by - trans B - · guard_target =ₛ A → B -- ensure we have `B` and not a free metavariable. - exact h - · guard_target =ₛ B → C - exact g - -set_option linter.unusedTactic false in -/-- `trans` for arrows between types. -/ -example {A B C : Type} (h : A → B) (g : B → C) : A → C := by - trans - guard_goal_nums 3 -- 3rd goal is the middle term - · exact h - · exact g - -universe u v w - -set_option linter.unusedTactic false in -/-- `trans` for arrows between types. -/ -example {A : Type u} {B : Type v} {C : Type w} (h : A → B) (g : B → C) : A → C := by - trans - guard_goal_nums 3 -- 3rd goal is the middle term - · exact h - · exact g From 9e3db2c4fc34b715fd1876983c6b12759a0ccb0d Mon Sep 17 00:00:00 2001 From: FR Date: Sat, 19 Oct 2024 12:38:54 +0000 Subject: [PATCH 180/505] perf: de-simp `tsub_eq_zero_of_le` (#17512) As mentioned [on Zulip](https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/Infrastructure.20for.20tracking.20frequently.20applied.20simp.20theorems/near/475331955): > `tsub_eq_zero_of_le` got worse in [the unmerged] #17444. It seems that Lean is searching `CanonicallyOrderedAdd` for types that are not ordered at all, then doing some slow unification of instances with metavariables that is impossible to succeed. --- Archive/Imo/Imo1986Q5.lean | 2 +- Mathlib/Algebra/MvPolynomial/Derivation.lean | 2 +- Mathlib/Algebra/Order/Sub/Basic.lean | 2 -- Mathlib/Data/ENNReal/Inv.lean | 2 +- Mathlib/Data/ENNReal/Operations.lean | 2 +- Mathlib/Data/Nat/Factorization/Basic.lean | 2 +- Mathlib/Data/Real/ConjExponents.lean | 4 ++-- Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean | 3 ++- Mathlib/MeasureTheory/Measure/DiracProba.lean | 2 +- Mathlib/Topology/Instances/ENNReal.lean | 2 +- Mathlib/Topology/Instances/ENat.lean | 2 +- Mathlib/Topology/MetricSpace/ThickenedIndicator.lean | 2 +- 12 files changed, 13 insertions(+), 14 deletions(-) diff --git a/Archive/Imo/Imo1986Q5.lean b/Archive/Imo/Imo1986Q5.lean index 82fe5961c1647..e8ca0c898f1d5 100644 --- a/Archive/Imo/Imo1986Q5.lean +++ b/Archive/Imo/Imo1986Q5.lean @@ -71,7 +71,7 @@ theorem isGood_iff {f : ℝ≥0 → ℝ≥0} : IsGood f ↔ f = fun x ↦ 2 / (2 refine ⟨fun hf ↦ funext hf.map_eq, ?_⟩ rintro rfl constructor - case map_two => simp + case map_two => simp [tsub_self] case map_ne_zero => intro x hx; simpa [tsub_eq_zero_iff_le] case map_add_rev => intro x y diff --git a/Mathlib/Algebra/MvPolynomial/Derivation.lean b/Mathlib/Algebra/MvPolynomial/Derivation.lean index 49d6860a55c77..1d88e3c8c50e2 100644 --- a/Mathlib/Algebra/MvPolynomial/Derivation.lean +++ b/Mathlib/Algebra/MvPolynomial/Derivation.lean @@ -45,7 +45,7 @@ theorem mkDerivationₗ_C (f : σ → A) (r : R) : mkDerivationₗ R f (C r) = 0 (mkDerivationₗ_monomial f _ _).trans (smul_zero _) theorem mkDerivationₗ_X (f : σ → A) (i : σ) : mkDerivationₗ R f (X i) = f i := - (mkDerivationₗ_monomial f _ _).trans <| by simp + (mkDerivationₗ_monomial f _ _).trans <| by simp [tsub_self] @[simp] theorem derivation_C (D : Derivation R (MvPolynomial σ R) A) (a : R) : D (C a) = 0 := diff --git a/Mathlib/Algebra/Order/Sub/Basic.lean b/Mathlib/Algebra/Order/Sub/Basic.lean index b2d6abd8dd435..e1b98ddf85486 100644 --- a/Mathlib/Algebra/Order/Sub/Basic.lean +++ b/Mathlib/Algebra/Order/Sub/Basic.lean @@ -32,8 +32,6 @@ theorem tsub_eq_zero_iff_le : a - b = 0 ↔ a ≤ b := by alias ⟨_, tsub_eq_zero_of_le⟩ := tsub_eq_zero_iff_le -attribute [simp] tsub_eq_zero_of_le - theorem tsub_self (a : α) : a - a = 0 := tsub_eq_zero_of_le le_rfl diff --git a/Mathlib/Data/ENNReal/Inv.lean b/Mathlib/Data/ENNReal/Inv.lean index 0f9c29a783d26..c0a3bde921b26 100644 --- a/Mathlib/Data/ENNReal/Inv.lean +++ b/Mathlib/Data/ENNReal/Inv.lean @@ -831,7 +831,7 @@ lemma smul_sSup {R} [SMul R ℝ≥0∞] [IsScalarTower R ℝ≥0∞ ℝ≥0∞] lemma sub_iSup [Nonempty ι] (ha : a ≠ ∞) : a - ⨆ i, f i = ⨅ i, a - f i := by obtain ⟨i, hi⟩ | h := em (∃ i, a < f i) · rw [tsub_eq_zero_iff_le.2 <| le_iSup_of_le _ hi.le, (iInf_eq_bot _).2, bot_eq_zero] - exact fun x hx ↦ ⟨i, by simpa [hi.le]⟩ + exact fun x hx ↦ ⟨i, by simpa [hi.le, tsub_eq_zero_of_le]⟩ simp_rw [not_exists, not_lt] at h refine le_antisymm (le_iInf fun i ↦ tsub_le_tsub_left (le_iSup ..) _) <| ENNReal.le_sub_of_add_le_left (ne_top_of_le_ne_top ha <| iSup_le h) <| diff --git a/Mathlib/Data/ENNReal/Operations.lean b/Mathlib/Data/ENNReal/Operations.lean index 26c974025ccf6..8318373b6a1bf 100644 --- a/Mathlib/Data/ENNReal/Operations.lean +++ b/Mathlib/Data/ENNReal/Operations.lean @@ -405,7 +405,7 @@ theorem sub_right_inj {a b c : ℝ≥0∞} (ha : a ≠ ∞) (hb : b ≤ a) (hc : (cancel_of_ne <| ne_top_of_le_ne_top ha hc) hb hc theorem sub_mul (h : 0 < b → b < a → c ≠ ∞) : (a - b) * c = a * c - b * c := by - rcases le_or_lt a b with hab | hab; · simp [hab, mul_right_mono hab] + rcases le_or_lt a b with hab | hab; · simp [hab, mul_right_mono hab, tsub_eq_zero_of_le] rcases eq_or_lt_of_le (zero_le b) with (rfl | hb); · simp exact (cancel_of_ne <| mul_ne_top hab.ne_top (h hb hab)).tsub_mul diff --git a/Mathlib/Data/Nat/Factorization/Basic.lean b/Mathlib/Data/Nat/Factorization/Basic.lean index 17af0fffcfa82..c8c06e616e2d7 100644 --- a/Mathlib/Data/Nat/Factorization/Basic.lean +++ b/Mathlib/Data/Nat/Factorization/Basic.lean @@ -187,7 +187,7 @@ theorem exists_factorization_lt_of_lt {a b : ℕ} (ha : a ≠ 0) (hab : a < b) : theorem factorization_div {d n : ℕ} (h : d ∣ n) : (n / d).factorization = n.factorization - d.factorization := by rcases eq_or_ne d 0 with (rfl | hd); · simp [zero_dvd_iff.mp h] - rcases eq_or_ne n 0 with (rfl | hn); · simp + rcases eq_or_ne n 0 with (rfl | hn); · simp [tsub_eq_zero_of_le] apply add_left_injective d.factorization simp only rw [tsub_add_cancel_of_le <| (Nat.factorization_le_iff_dvd hd hn).mpr h, ← diff --git a/Mathlib/Data/Real/ConjExponents.lean b/Mathlib/Data/Real/ConjExponents.lean index 04bfc20da0ec1..581ac0eb8bed3 100644 --- a/Mathlib/Data/Real/ConjExponents.lean +++ b/Mathlib/Data/Real/ConjExponents.lean @@ -272,7 +272,7 @@ protected lemma conjExponent (hp : 1 ≤ p) : p.IsConjExponent (conjExponent p) refine (AddLECancellable.eq_tsub_iff_add_eq_of_le (α := ℝ≥0∞) (by simpa) (by simpa)).1 ?_ rw [inv_eq_iff_eq_inv] obtain rfl | hp₁ := hp.eq_or_lt - · simp + · simp [tsub_eq_zero_of_le] obtain rfl | hp := eq_or_ne p ∞ · simp calc @@ -317,7 +317,7 @@ lemma mul_eq_add : p * q = p + q := by lemma div_conj_eq_sub_one : p / q = p - 1 := by obtain rfl | hq := eq_or_ne q ∞ - · simp [h.symm.conj_eq] + · simp [h.symm.conj_eq, tsub_eq_zero_of_le] refine ENNReal.eq_sub_of_add_eq one_ne_top ?_ rw [← ENNReal.div_self h.symm.ne_zero hq, ← ENNReal.add_div, ← h.mul_eq_add, mul_div_assoc, ENNReal.div_self h.symm.ne_zero hq, mul_one] diff --git a/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean b/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean index 9dfc57bb16291..ed5e038f6f5c4 100644 --- a/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean +++ b/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean @@ -255,7 +255,8 @@ instance instMeasurableMul₂ : MeasurableMul₂ ℝ≥0∞ := by instance instMeasurableSub₂ : MeasurableSub₂ ℝ≥0∞ := ⟨by apply measurable_of_measurable_nnreal_nnreal <;> - simp [← WithTop.coe_sub]; exact continuous_sub.measurable.coe_nnreal_ennreal⟩ + simp [← WithTop.coe_sub, tsub_eq_zero_of_le]; + exact continuous_sub.measurable.coe_nnreal_ennreal⟩ instance instMeasurableInv : MeasurableInv ℝ≥0∞ := ⟨continuous_inv.measurable⟩ diff --git a/Mathlib/MeasureTheory/Measure/DiracProba.lean b/Mathlib/MeasureTheory/Measure/DiracProba.lean index 18bc9a7f40a5e..51377c0135958 100644 --- a/Mathlib/MeasureTheory/Measure/DiracProba.lean +++ b/Mathlib/MeasureTheory/Measure/DiracProba.lean @@ -36,7 +36,7 @@ lemma CompletelyRegularSpace.exists_BCNN {X : Type*} [TopologicalSpace X] [Compl set g' := BoundedContinuousFunction.mkOfBound ⟨fun x ↦ Real.toNNReal (g x), continuous_real_toNNReal.comp g_cont.subtype_val⟩ 1 g_bdd set f := 1 - g' - refine ⟨f, by simp [f, g', gx_zero], fun y y_in_K ↦ by simp [f, g', g_one_on_K y_in_K]⟩ + refine ⟨f, by simp [f, g', gx_zero], fun y y_in_K ↦ by simp [f, g', g_one_on_K y_in_K, tsub_self]⟩ namespace MeasureTheory diff --git a/Mathlib/Topology/Instances/ENNReal.lean b/Mathlib/Topology/Instances/ENNReal.lean index a6719468d25d3..f5c72e34f691d 100644 --- a/Mathlib/Topology/Instances/ENNReal.lean +++ b/Mathlib/Topology/Instances/ENNReal.lean @@ -422,7 +422,7 @@ theorem continuousOn_sub_left (a : ℝ≥0∞) : ContinuousOn (a - ·) { x : ℝ theorem continuous_sub_right (a : ℝ≥0∞) : Continuous fun x : ℝ≥0∞ => x - a := by by_cases a_infty : a = ∞ - · simp [a_infty, continuous_const] + · simp [a_infty, continuous_const, tsub_eq_zero_of_le] · rw [show (fun x => x - a) = (fun p : ℝ≥0∞ × ℝ≥0∞ => p.fst - p.snd) ∘ fun x => ⟨x, a⟩ by rfl] apply ContinuousOn.comp_continuous continuousOn_sub (continuous_id'.prod_mk continuous_const) intro x diff --git a/Mathlib/Topology/Instances/ENat.lean b/Mathlib/Topology/Instances/ENat.lean index 0217f6abc247c..c69882b280e04 100644 --- a/Mathlib/Topology/Instances/ENat.lean +++ b/Mathlib/Topology/Instances/ENat.lean @@ -92,7 +92,7 @@ protected theorem continuousAt_sub {a b : ℕ∞} (h : a ≠ ⊤ ∨ b ≠ ⊤) simpa [ContinuousAt, nhds_prod_eq] using tendsto_pure_nhds _ _ | (a : ℕ), ⊤, _ => suffices ∀ᶠ b in 𝓝 ⊤, (a - b : ℕ∞) = 0 by - simpa [ContinuousAt, nhds_prod_eq] + simpa [ContinuousAt, nhds_prod_eq, tsub_eq_zero_of_le] filter_upwards [le_mem_nhds (WithTop.coe_lt_top a)] with b using tsub_eq_zero_of_le | ⊤, (b : ℕ), _ => suffices ∀ n : ℕ, ∀ᶠ a : ℕ∞ in 𝓝 ⊤, b + n < a by diff --git a/Mathlib/Topology/MetricSpace/ThickenedIndicator.lean b/Mathlib/Topology/MetricSpace/ThickenedIndicator.lean index 50b3c349f536c..1df15588c66b4 100644 --- a/Mathlib/Topology/MetricSpace/ThickenedIndicator.lean +++ b/Mathlib/Topology/MetricSpace/ThickenedIndicator.lean @@ -86,7 +86,7 @@ theorem thickenedIndicatorAux_zero {δ : ℝ} (δ_pos : 0 < δ) (E : Set α) {x have key := tsub_le_tsub (@rfl _ (1 : ℝ≥0∞)).le (ENNReal.div_le_div x_out (@rfl _ (ENNReal.ofReal δ : ℝ≥0∞)).le) rw [ENNReal.div_self (ne_of_gt (ENNReal.ofReal_pos.mpr δ_pos)) ofReal_ne_top] at key - simpa using key + simpa [tsub_self] using key theorem thickenedIndicatorAux_mono {δ₁ δ₂ : ℝ} (hle : δ₁ ≤ δ₂) (E : Set α) : thickenedIndicatorAux δ₁ E ≤ thickenedIndicatorAux δ₂ E := From 2b96a2f7f0879e2bd4f2517079e0752c24206277 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Calle=20S=C3=B6nne?= Date: Sat, 19 Oct 2024 13:30:30 +0000 Subject: [PATCH 181/505] feat(CategoryTheory/Bicategory): Grothendieck construction for a pseudofunctor (#15419) In this PR we add the Grothendieck construction for pseudofunctors. --- Mathlib.lean | 1 + .../Bicategory/Grothendieck.lean | 122 ++++++++++++++++++ docs/references.bib | 10 ++ 3 files changed, 133 insertions(+) create mode 100644 Mathlib/CategoryTheory/Bicategory/Grothendieck.lean diff --git a/Mathlib.lean b/Mathlib.lean index e24762ba7699b..f51a308d8a11e 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1462,6 +1462,7 @@ import Mathlib.CategoryTheory.Bicategory.Functor.Oplax import Mathlib.CategoryTheory.Bicategory.Functor.Prelax import Mathlib.CategoryTheory.Bicategory.Functor.Pseudofunctor import Mathlib.CategoryTheory.Bicategory.FunctorBicategory +import Mathlib.CategoryTheory.Bicategory.Grothendieck import Mathlib.CategoryTheory.Bicategory.Kan.Adjunction import Mathlib.CategoryTheory.Bicategory.Kan.HasKan import Mathlib.CategoryTheory.Bicategory.Kan.IsKan diff --git a/Mathlib/CategoryTheory/Bicategory/Grothendieck.lean b/Mathlib/CategoryTheory/Bicategory/Grothendieck.lean new file mode 100644 index 0000000000000..d7eb112f1f6ba --- /dev/null +++ b/Mathlib/CategoryTheory/Bicategory/Grothendieck.lean @@ -0,0 +1,122 @@ +/- +Copyright (c) 2024 Calle Sönne. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Calle Sönne +-/ + +import Mathlib.CategoryTheory.Bicategory.LocallyDiscrete + +/-! +# The Grothendieck construction + +Given a category `𝒮` and any pseudofunctor `F` from `𝒮ᵒᵖ` to `Cat`, we associate to it a category +`∫ F`, equipped with a functor `∫ F ⥤ 𝒮`. + +The category `∫ F` is defined as follows: +* Objects: pairs `(S, a)` where `S` is an object of the base category and `a` is an object of the + category `F(S)`. +* Morphisms: morphisms `(R, b) ⟶ (S, a)` are defined as pairs `(f, h)` where `f : R ⟶ S` is a + morphism in `𝒮` and `h : b ⟶ F(f)(a)` + +The projection functor `∫ F ⥤ 𝒮` is then given by projecting to the first factors, i.e. +* On objects, it sends `(S, a)` to `S` +* On morphisms, it sends `(f, h)` to `f` + +## References +[Vistoli2008] "Notes on Grothendieck Topologies, Fibered Categories and Descent Theory" by +Angelo Vistoli +-/ + +namespace CategoryTheory + +universe w v₁ v₂ v₃ u₁ u₂ u₃ + +open CategoryTheory Functor Category Opposite Discrete Bicategory + +variable {𝒮 : Type u₁} [Category.{v₁} 𝒮] {F : Pseudofunctor (LocallyDiscrete 𝒮ᵒᵖ) Cat.{v₂, u₂}} + +/-- The type of objects in the fibered category associated to a presheaf valued in types. -/ +@[ext] +structure Pseudofunctor.Grothendieck (F : Pseudofunctor (LocallyDiscrete 𝒮ᵒᵖ) Cat.{v₂, u₂}) where + /-- The underlying object in the base category. -/ + base : 𝒮 + /-- The object in the fiber of the base object. -/ + fiber : F.obj ⟨op base⟩ + +namespace Pseudofunctor.Grothendieck + +/-- Notation for the Grothendieck category associated to a pseudofunctor `F`. -/ +scoped prefix:75 "∫ " => Pseudofunctor.Grothendieck + +/-- A morphism in the Grothendieck category `F : C ⥤ Cat` consists of +`base : X.base ⟶ Y.base` and `f.fiber : (F.map base).obj X.fiber ⟶ Y.fiber`. +-/ +structure Hom (X Y : ∫ F) where + /-- The morphism between base objects. -/ + base : X.base ⟶ Y.base + /-- The morphism in the fiber over the domain. -/ + fiber : X.fiber ⟶ (F.map base.op.toLoc).obj Y.fiber + +@[simps!] +instance categoryStruct : CategoryStruct (∫ F) where + Hom X Y := Hom X Y + id X := { + base := 𝟙 X.base + fiber := (F.mapId ⟨op X.base⟩).inv.app X.fiber } + comp {_ _ Z} f g := { + base := f.base ≫ g.base + fiber := f.fiber ≫ (F.map f.base.op.toLoc).map g.fiber ≫ + (F.mapComp g.base.op.toLoc f.base.op.toLoc).inv.app Z.fiber } + +section + +variable {a b : ∫ F} + +@[ext (iff := false)] +lemma Hom.ext (f g : a ⟶ b) (hfg₁ : f.base = g.base) + (hfg₂ : f.fiber = g.fiber ≫ eqToHom (hfg₁ ▸ rfl)) : f = g := by + cases f; cases g + congr + dsimp at hfg₁ + rw [← conj_eqToHom_iff_heq _ _ rfl (hfg₁ ▸ rfl)] + simpa only [eqToHom_refl, id_comp] using hfg₂ + +lemma Hom.ext_iff (f g : a ⟶ b) : + f = g ↔ ∃ (hfg : f.base = g.base), f.fiber = g.fiber ≫ eqToHom (hfg ▸ rfl) where + mp hfg := ⟨by rw [hfg], by simp [hfg]⟩ + mpr := fun ⟨hfg₁, hfg₂⟩ => Hom.ext f g hfg₁ hfg₂ + +lemma Hom.congr {a b : ∫ F} {f g : a ⟶ b} (h : f = g) : + f.fiber = g.fiber ≫ eqToHom (h ▸ rfl) := by + simp [h] + +end + +/-- The category structure on `∫ F`. -/ +instance category : Category (∫ F) where + toCategoryStruct := Pseudofunctor.Grothendieck.categoryStruct + id_comp {a b} f := by + ext + · simp + · simp [F.mapComp_id_right_inv_app, Strict.rightUnitor_eqToIso, ← NatTrans.naturality_assoc] + comp_id {a b} f := by + ext + · simp + · simp [F.mapComp_id_left_inv_app, ← Functor.map_comp_assoc, Strict.leftUnitor_eqToIso] + assoc f g h := by + ext + · simp + · simp [← NatTrans.naturality_assoc, F.mapComp_assoc_right_inv_app, Strict.associator_eqToIso] + +variable (F) + +/-- The projection `∫ F ⥤ 𝒮` given by projecting both objects and homs to the first +factor. -/ +@[simps] +def forget : ∫ F ⥤ 𝒮 where + obj X := X.base + map f := f.base + +end Pseudofunctor.Grothendieck + +end CategoryTheory diff --git a/docs/references.bib b/docs/references.bib index 39e500681a649..e94cd696fd232 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -3287,6 +3287,16 @@ @Book{ verdier1996 mrnumber = {1453167} } +@Misc{ vistoli2004, + author = {Vistoli, Angelo}, + title = {Notes on {Grothendieck} topologies, fibered categories and + descent theory}, + year = {2004}, + howpublished = {Preprint, {arXiv}:math/0412512 [math.{AG}] (2004)}, + url = {https://arxiv.org/abs/math/0412512}, + arxiv = {arXiv:math/0412512} +} + @Book{ wall2018analytic, title = {Analytic Theory of Continued Fractions}, author = {Wall, H.S.}, From 92412db77694c08a00efa5b1028264e8c7079f7b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Calle=20S=C3=B6nne?= Date: Sat, 19 Oct 2024 13:51:18 +0000 Subject: [PATCH 182/505] feat(FiberedCategory/Fiber): define the fibers of a functor (#13603) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We define the fibers of a functor `p : 𝒳 ⥤ 𝒮` over objects `S : 𝒮`, and develop some basic API. Relatively little API is developed as in an upcoming PR we introduce a class `HasFibers` which should be preferred. Co-authored-by: Paul Lezeau --- Mathlib.lean | 1 + .../CategoryTheory/FiberedCategory/Fiber.lean | 127 ++++++++++++++++++ .../FiberedCategory/HomLift.lean | 10 +- 3 files changed, 135 insertions(+), 3 deletions(-) create mode 100644 Mathlib/CategoryTheory/FiberedCategory/Fiber.lean diff --git a/Mathlib.lean b/Mathlib.lean index f51a308d8a11e..fbd2d2cec89b5 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1547,6 +1547,7 @@ import Mathlib.CategoryTheory.Extensive import Mathlib.CategoryTheory.FiberedCategory.BasedCategory import Mathlib.CategoryTheory.FiberedCategory.Cartesian import Mathlib.CategoryTheory.FiberedCategory.Cocartesian +import Mathlib.CategoryTheory.FiberedCategory.Fiber import Mathlib.CategoryTheory.FiberedCategory.Fibered import Mathlib.CategoryTheory.FiberedCategory.HomLift import Mathlib.CategoryTheory.Filtered.Basic diff --git a/Mathlib/CategoryTheory/FiberedCategory/Fiber.lean b/Mathlib/CategoryTheory/FiberedCategory/Fiber.lean new file mode 100644 index 0000000000000..94cfb1727d195 --- /dev/null +++ b/Mathlib/CategoryTheory/FiberedCategory/Fiber.lean @@ -0,0 +1,127 @@ +/- +Copyright (c) 2024 Calle Sönne. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Calle Sönne, Paul Lezeau +-/ + +import Mathlib.CategoryTheory.FiberedCategory.HomLift +import Mathlib.CategoryTheory.Functor.Const + +/-! + +# Fibers of a functors + +In this file we define, for a functor `p : 𝒳 ⥤ 𝒴`, the fiber categories `Fiber p S` for every +`S : 𝒮` as follows +- An object in `Fiber p S` is a pair `(a, ha)` where `a : 𝒳` and `ha : p.obj a = S`. +- A morphism in `Fiber p S` is a morphism `φ : a ⟶ b` in 𝒳 such that `p.map φ = 𝟙 S`. + +For any category `C` equipped with a functor `F : C ⥤ 𝒳` such that `F ⋙ p` is constant at `S`, +we define a functor `inducedFunctor : C ⥤ Fiber p S` that `F` factors through. +-/ + +universe v₁ u₁ v₂ u₂ v₃ u₃ + +namespace CategoryTheory + +open IsHomLift + +namespace Functor + +variable {𝒮 : Type u₁} {𝒳 : Type u₂} [Category.{v₁} 𝒮] [Category.{v₂} 𝒳] + +/-- `Fiber p S` is the type of elements of `𝒳` mapping to `S` via `p`. -/ +def Fiber (p : 𝒳 ⥤ 𝒮) (S : 𝒮) := { a : 𝒳 // p.obj a = S } + +namespace Fiber + +variable {p : 𝒳 ⥤ 𝒮} {S : 𝒮} + +/-- `Fiber p S` has the structure of a category with morphisms being those lying over `𝟙 S`. -/ +instance fiberCategory : Category (Fiber p S) where + Hom a b := {φ : a.1 ⟶ b.1 // IsHomLift p (𝟙 S) φ} + id a := ⟨𝟙 a.1, IsHomLift.id a.2⟩ + comp φ ψ := ⟨φ.val ≫ ψ.val, by have := φ.2; have := ψ.2; infer_instance⟩ + +/-- The functor including `Fiber p S` into `𝒳`. -/ +def fiberInclusion : Fiber p S ⥤ 𝒳 where + obj a := a.1 + map φ := φ.1 + +instance {a b : Fiber p S} (φ : a ⟶ b) : IsHomLift p (𝟙 S) (fiberInclusion.map φ) := φ.2 + +@[ext] +lemma hom_ext {a b : Fiber p S} {φ ψ : a ⟶ b} + (h : fiberInclusion.map φ = fiberInclusion.map ψ) : φ = ψ := + Subtype.ext h + +instance : (fiberInclusion : Fiber p S ⥤ _).Faithful where + +/-- For fixed `S : 𝒮` this is the natural isomorphism between `fiberInclusion ⋙ p` and the constant +function valued at `S`. -/ +@[simps!] +def fiberInclusionCompIsoConst : fiberInclusion ⋙ p ≅ (const (Fiber p S)).obj S := + NatIso.ofComponents (fun X ↦ eqToIso X.2) + (fun φ ↦ by simp [IsHomLift.fac' p (𝟙 S) (fiberInclusion.map φ)]) + +lemma fiberInclusion_comp_eq_const : fiberInclusion ⋙ p = (const (Fiber p S)).obj S := + Functor.ext (fun x ↦ x.2) (fun _ _ φ ↦ IsHomLift.fac' p (𝟙 S) (fiberInclusion.map φ)) + +/-- The object of the fiber over `S` corresponding to a `a : 𝒳` such that `p(a) = S`. -/ +def mk {p : 𝒳 ⥤ 𝒮} {S : 𝒮} {a : 𝒳} (ha : p.obj a = S) : Fiber p S := ⟨a, ha⟩ + +@[simp] +lemma fiberInclusion_mk {p : 𝒳 ⥤ 𝒮} {S : 𝒮} {a : 𝒳} (ha : p.obj a = S) : + fiberInclusion.obj (mk ha) = a := + rfl + +/-- The morphism in the fiber over `S` corresponding to a morphism in `𝒳` lifting `𝟙 S`. -/ +def homMk (p : 𝒳 ⥤ 𝒮) (S : 𝒮) {a b : 𝒳} (φ : a ⟶ b) [IsHomLift p (𝟙 S) φ] : + mk (domain_eq p (𝟙 S) φ) ⟶ mk (codomain_eq p (𝟙 S) φ) := + ⟨φ, inferInstance⟩ + +@[simp] +lemma fiberInclusion_homMk (p : 𝒳 ⥤ 𝒮) (S : 𝒮) {a b : 𝒳} (φ : a ⟶ b) [IsHomLift p (𝟙 S) φ] : + fiberInclusion.map (homMk p S φ) = φ := + rfl + +@[simp] +lemma homMk_id (p : 𝒳 ⥤ 𝒮) (S : 𝒮) (a : 𝒳) [IsHomLift p (𝟙 S) (𝟙 a)] : + homMk p S (𝟙 a) = 𝟙 (mk (domain_eq p (𝟙 S) (𝟙 a))) := + rfl + +@[simp] +lemma homMk_comp {a b c : 𝒳} (φ : a ⟶ b) (ψ : b ⟶ c) [IsHomLift p (𝟙 S) φ] + [IsHomLift p (𝟙 S) ψ] : homMk p S φ ≫ homMk p S ψ = homMk p S (φ ≫ ψ) := + rfl + +section + +variable {p : 𝒳 ⥤ 𝒮} {S : 𝒮} {C : Type u₃} [Category.{v₃} C] {F : C ⥤ 𝒳} + (hF : F ⋙ p = (const C).obj S) + +/-- Given a functor `F : C ⥤ 𝒳` such that `F ⋙ p` is constant at some `S : 𝒮`, then +we get an induced functor `C ⥤ Fiber p S` that `F` factors through. -/ +@[simps] +def inducedFunctor : C ⥤ Fiber p S where + obj x := ⟨F.obj x, by simp only [← comp_obj, hF, const_obj_obj]⟩ + map φ := ⟨F.map φ, of_commsq _ _ _ _ _ <| by simpa using (eqToIso hF).hom.naturality φ⟩ + +@[simp] +lemma inducedFunctor_map {X Y : C} (f : X ⟶ Y) : + fiberInclusion.map ((inducedFunctor hF).map f) = F.map f := rfl + +/-- Given a functor `F : C ⥤ 𝒳` such that `F ⋙ p` is constant at some `S : 𝒮`, then +we get a natural isomorphism between `inducedFunctor _ ⋙ fiberInclusion` and `F`. -/ +@[simps!] +def inducedFunctorCompIsoSelf : (inducedFunctor hF) ⋙ fiberInclusion ≅ F := Iso.refl _ + +lemma inducedFunctor_comp : (inducedFunctor hF) ⋙ fiberInclusion = F := rfl + +end + +end Fiber + +end Functor + +end CategoryTheory diff --git a/Mathlib/CategoryTheory/FiberedCategory/HomLift.lean b/Mathlib/CategoryTheory/FiberedCategory/HomLift.lean index b9bc045f14acb..5d64e6e16c258 100644 --- a/Mathlib/CategoryTheory/FiberedCategory/HomLift.lean +++ b/Mathlib/CategoryTheory/FiberedCategory/HomLift.lean @@ -106,12 +106,16 @@ lemma of_fac' {R S : 𝒮} {a b : 𝒳} (f : R ⟶ S) (φ : a ⟶ b) (ha : p.obj obtain rfl : f = p.map φ := by simpa using h.symm infer_instance -lemma of_commSq {R S : 𝒮} {a b : 𝒳} (f : R ⟶ S) (φ : a ⟶ b) (ha : p.obj a = R) (hb : p.obj b = S) - (h : CommSq (p.map φ) (eqToHom ha) (eqToHom hb) f) : p.IsHomLift f φ := by +lemma of_commsq {R S : 𝒮} {a b : 𝒳} (f : R ⟶ S) (φ : a ⟶ b) (ha : p.obj a = R) (hb : p.obj b = S) + (h : p.map φ ≫ eqToHom hb = (eqToHom ha) ≫ f) : p.IsHomLift f φ := by subst ha hb - obtain rfl : f = p.map φ := by simpa using h.1.symm + obtain rfl : f = p.map φ := by simpa using h.symm infer_instance +lemma of_commSq {R S : 𝒮} {a b : 𝒳} (f : R ⟶ S) (φ : a ⟶ b) (ha : p.obj a = R) (hb : p.obj b = S) + (h : CommSq (p.map φ) (eqToHom ha) (eqToHom hb) f) : p.IsHomLift f φ := + of_commsq p f φ ha hb h.1 + instance comp {R S T : 𝒮} {a b c : 𝒳} (f : R ⟶ S) (g : S ⟶ T) (φ : a ⟶ b) (ψ : b ⟶ c) [p.IsHomLift f φ] [p.IsHomLift g ψ] : p.IsHomLift (f ≫ g) (φ ≫ ψ) := by apply of_commSq From 0d932abadf1d7aad0fd9efbae434ff5e6fa24667 Mon Sep 17 00:00:00 2001 From: Daniel Weber Date: Sat, 19 Oct 2024 14:07:56 +0000 Subject: [PATCH 183/505] feat: define rooted trees (#16272) This is preparation for proving Kruskal's tree theorem. Co-authored-by: Daniel Weber <55664973+Command-Master@users.noreply.github.com> --- Mathlib.lean | 1 + Mathlib/Order/SuccPred/Tree.lean | 228 +++++++++++++++++++++++++++++++ 2 files changed, 229 insertions(+) create mode 100644 Mathlib/Order/SuccPred/Tree.lean diff --git a/Mathlib.lean b/Mathlib.lean index fbd2d2cec89b5..4c5d00a1fa732 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3797,6 +3797,7 @@ import Mathlib.Order.SuccPred.IntervalSucc import Mathlib.Order.SuccPred.Limit import Mathlib.Order.SuccPred.LinearLocallyFinite import Mathlib.Order.SuccPred.Relation +import Mathlib.Order.SuccPred.Tree import Mathlib.Order.SupClosed import Mathlib.Order.SupIndep import Mathlib.Order.SymmDiff diff --git a/Mathlib/Order/SuccPred/Tree.lean b/Mathlib/Order/SuccPred/Tree.lean new file mode 100644 index 0000000000000..ea7ab54e7fb12 --- /dev/null +++ b/Mathlib/Order/SuccPred/Tree.lean @@ -0,0 +1,228 @@ +/- +Copyright (c) 2024 Daniel Weber. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Daniel Weber +-/ + +import Mathlib.Order.SuccPred.Archimedean +import Mathlib.Data.Nat.Find +import Mathlib.Order.Atoms +import Mathlib.Data.SetLike.Basic + +/-! +# Rooted trees + +This file proves basic results about rooted trees, represented using the ancestorship order. +This is a `PartialOrder`, with `PredOrder` with the immediate parent as a predecessor, and an +`OrderBot` which is the root. We also have an `IsPredArchimedean` assumption to prevent infinite +dangling chains. + +--/ + +variable {α : Type*} [PartialOrder α] [PredOrder α] [IsPredArchimedean α] + +namespace IsPredArchimedean + +variable [OrderBot α] + +section DecidableEq + +variable [DecidableEq α] + +/-- +The unique atom less than an element in an `OrderBot` with archimedean predecessor. +-/ +def findAtom (r : α) : α := + Order.pred^[Nat.find (bot_le (a := r)).exists_pred_iterate - 1] r + +@[simp] +lemma findAtom_le (r : α) : findAtom r ≤ r := + Order.pred_iterate_le _ _ + +@[simp] +lemma findAtom_bot : findAtom (⊥ : α) = ⊥ := by + apply Function.iterate_fixed + simp + +@[simp] +lemma pred_findAtom (r : α) : Order.pred (findAtom r) = ⊥ := by + unfold findAtom + generalize h : Nat.find (bot_le (a := r)).exists_pred_iterate = n + cases n + · have : Order.pred^[0] r = ⊥ := by + rw [← h] + apply Nat.find_spec (bot_le (a := r)).exists_pred_iterate + simp only [Function.iterate_zero, id_eq] at this + simp [this] + · simp only [Nat.add_sub_cancel_right, ← Function.iterate_succ_apply', Nat.succ_eq_add_one] + rw [← h] + apply Nat.find_spec (bot_le (a := r)).exists_pred_iterate + +@[simp] +lemma findAtom_eq_bot {r : α} : + findAtom r = ⊥ ↔ r = ⊥ where + mp h := by + unfold findAtom at h + have := Nat.find_min' (bot_le (a := r)).exists_pred_iterate h + replace : Nat.find (bot_le (a := r)).exists_pred_iterate = 0 := by omega + simpa [this] using h + mpr h := by simp [h] + +lemma findAtom_ne_bot {r : α} : + findAtom r ≠ ⊥ ↔ r ≠ ⊥ := findAtom_eq_bot.not + +lemma isAtom_findAtom {r : α} (hr : r ≠ ⊥) : + IsAtom (findAtom r) := by + constructor + · simp [hr] + · intro b hb + apply Order.le_pred_of_lt at hb + simpa using hb + +@[simp] +lemma isAtom_findAtom_iff {r : α} : + IsAtom (findAtom r) ↔ r ≠ ⊥ where + mpr := isAtom_findAtom + mp h nh := by simp only [nh, findAtom_bot] at h; exact h.1 rfl + +end DecidableEq + +instance instIsAtomic : IsAtomic α where + eq_bot_or_exists_atom_le b := by classical + rw [or_iff_not_imp_left] + intro hb + use findAtom b, isAtom_findAtom hb, findAtom_le b + +end IsPredArchimedean + +/-- +The type of rooted trees. +-/ +structure RootedTree where + /-- The type representing the elements in the tree. -/ + α : Type* + /-- The type should be a `SemilatticeInf`, + where `inf` is the least common ancestor in the tree. -/ + [semilatticeInf : SemilatticeInf α] + /-- The type should have a bottom, the root. -/ + [orderBot : OrderBot α] + /-- The type should have a predecessor for every element, its parent. -/ + [predOrder : PredOrder α] + /-- The predecessor relationship should be archimedean. -/ + [isPredArchimedean : IsPredArchimedean α] + +attribute [coe] RootedTree.α + +instance : CoeSort RootedTree Type* := ⟨RootedTree.α⟩ + +attribute [instance] RootedTree.semilatticeInf RootedTree.predOrder + RootedTree.orderBot RootedTree.isPredArchimedean + +/-- +A subtree is represented by its root, therefore this is a type synonym. +-/ +def SubRootedTree (t : RootedTree) : Type* := t + +/-- +The root of a `SubRootedTree`. +-/ +def SubRootedTree.root {t : RootedTree} (v : SubRootedTree t) : t := v + +/-- +The `SubRootedTree` rooted at a given node. +-/ +def RootedTree.subtree (t : RootedTree) (r : t) : SubRootedTree t := r + +@[simp] +lemma RootedTree.root_subtree (t : RootedTree) (r : t) : (t.subtree r).root = r := rfl + +@[simp] +lemma RootedTree.subtree_root (t : RootedTree) (v : SubRootedTree t) : t.subtree v.root = v := rfl + +@[ext] +lemma SubRootedTree.ext {t : RootedTree} {v₁ v₂ : SubRootedTree t} + (h : v₁.root = v₂.root) : v₁ = v₂ := h + +instance (t : RootedTree) : SetLike (SubRootedTree t) t where + coe v := Set.Ici v.root + coe_injective' a₁ a₂ h := by + simpa only [Set.Ici_inj, ← SubRootedTree.ext_iff] using h + +lemma SubRootedTree.mem_iff {t : RootedTree} {r : SubRootedTree t} {v : t} : + v ∈ r ↔ r.root ≤ v := Iff.rfl + +/-- +The coercion from a `SubRootedTree` to a `RootedTree`. +-/ +@[coe, reducible] +noncomputable def SubRootedTree.coeTree {t : RootedTree} (r : SubRootedTree t) : RootedTree where + α := Set.Ici r.root + +noncomputable instance (t : RootedTree) : CoeOut (SubRootedTree t) RootedTree := + ⟨SubRootedTree.coeTree⟩ + +@[simp] +lemma SubRootedTree.bot_mem_iff {t : RootedTree} (r : SubRootedTree t) : + ⊥ ∈ r ↔ r.root = ⊥ := by + simp [mem_iff] + +/-- +All of the immediate subtrees of a given rooted tree, that is subtrees which are rooted at a direct +child of the root (or, order theoretically, at an atom). +-/ +def RootedTree.subtrees (t : RootedTree) : Set (SubRootedTree t) := + {x | IsAtom x.root} + +variable {t : RootedTree} + +lemma SubRootedTree.root_ne_bot_of_mem_subtrees (r : SubRootedTree t) (hr : r ∈ t.subtrees) : + r.root ≠ ⊥ := by + simp only [RootedTree.subtrees, Set.mem_setOf_eq] at hr + exact hr.1 + +lemma RootedTree.mem_subtrees_disjoint_iff {t₁ t₂ : SubRootedTree t} + (ht₁ : t₁ ∈ t.subtrees) (ht₂ : t₂ ∈ t.subtrees) (v₁ v₂ : t) (h₁ : v₁ ∈ t₁) + (h₂ : v₂ ∈ t₂) : + Disjoint v₁ v₂ ↔ t₁ ≠ t₂ where + mp h := by + intro nh + have : t₁.root ≤ (v₁ : t) ⊓ (v₂ : t) := by + simp only [le_inf_iff] + exact ⟨h₁, nh ▸ h₂⟩ + rw [h.eq_bot] at this + simp only [le_bot_iff] at this + exact t₁.root_ne_bot_of_mem_subtrees ht₁ this + mpr h := by + rw [SubRootedTree.mem_iff] at h₁ h₂ + contrapose! h + rw [disjoint_iff, ← ne_eq, ← bot_lt_iff_ne_bot] at h + rcases lt_or_le_of_directed (by simp : v₁ ⊓ v₂ ≤ v₁) h₁ with oh | oh + · simp_all [RootedTree.subtrees, IsAtom.lt_iff] + rw [le_inf_iff] at oh + ext + simpa only [ht₂.le_iff_eq ht₁.1, ht₁.le_iff_eq ht₂.1, eq_comm, or_self] using + le_total_of_directed oh.2 h₂ + +lemma RootedTree.subtrees_disjoint : t.subtrees.PairwiseDisjoint ((↑) : _ → Set t) := by + intro t₁ ht₁ t₂ ht₂ h + rw [Function.onFun_apply, Set.disjoint_left] + intro a ha hb + rw [← mem_subtrees_disjoint_iff ht₁ ht₂ a a ha hb, disjoint_self] at h + subst h + simp only [SetLike.mem_coe, SubRootedTree.bot_mem_iff] at ha + exact t₁.root_ne_bot_of_mem_subtrees ht₁ ha + +/-- +The immediate subtree of `t` containing `v`, or all of `t` if `v` is the root. +-/ +def RootedTree.subtreeOf (t : RootedTree) [DecidableEq t] (v : t) : SubRootedTree t := + t.subtree (IsPredArchimedean.findAtom v) + +@[simp] +lemma RootedTree.mem_subtreeOf [DecidableEq t] {v : t} : + v ∈ t.subtreeOf v := by + simp [SubRootedTree.mem_iff, RootedTree.subtreeOf] + +lemma RootedTree.subtreeOf_mem_subtrees [DecidableEq t] {v : t} (hr : v ≠ ⊥) : + t.subtreeOf v ∈ t.subtrees := by + simpa [RootedTree.subtrees, RootedTree.subtreeOf] From c8922d33629788528cbfa02f9c56afb909d1c72d Mon Sep 17 00:00:00 2001 From: FR Date: Sat, 19 Oct 2024 15:25:40 +0000 Subject: [PATCH 184/505] chore: make `Quotient.mk''` an abbrev of `Quotient.mk _` (#16264) The plan is to deprecate `Quotient.mk''` in future PRs. --- Mathlib/Algebra/BigOperators/Group/Finset.lean | 4 ++-- Mathlib/Algebra/CharZero/Quotient.lean | 6 +++--- Mathlib/Data/List/Cycle.lean | 6 +++--- Mathlib/Data/Multiset/Basic.lean | 3 +++ Mathlib/Data/Quot.lean | 6 ++---- Mathlib/Data/Set/Image.lean | 2 +- Mathlib/GroupTheory/Coset/Basic.lean | 4 ++-- Mathlib/GroupTheory/GroupAction/Quotient.lean | 6 +++--- Mathlib/GroupTheory/PGroup.lean | 2 +- 9 files changed, 20 insertions(+), 19 deletions(-) diff --git a/Mathlib/Algebra/BigOperators/Group/Finset.lean b/Mathlib/Algebra/BigOperators/Group/Finset.lean index ed79af0319d28..1dbe92862bdb6 100644 --- a/Mathlib/Algebra/BigOperators/Group/Finset.lean +++ b/Mathlib/Algebra/BigOperators/Group/Finset.lean @@ -1611,13 +1611,13 @@ theorem dvd_prod_of_mem (f : α → β) {a : α} {s : Finset α} (ha : a ∈ s) /-- A product can be partitioned into a product of products, each equivalent under a setoid. -/ @[to_additive "A sum can be partitioned into a sum of sums, each equivalent under a setoid."] theorem prod_partition (R : Setoid α) [DecidableRel R.r] : - ∏ x ∈ s, f x = ∏ xbar ∈ s.image Quotient.mk'', ∏ y ∈ s.filter (⟦·⟧ = xbar), f y := by + ∏ x ∈ s, f x = ∏ xbar ∈ s.image (Quotient.mk _), ∏ y ∈ s.filter (⟦·⟧ = xbar), f y := by refine (Finset.prod_image' f fun x _hx => ?_).symm rfl /-- If we can partition a product into subsets that cancel out, then the whole product cancels. -/ @[to_additive "If we can partition a sum into subsets that cancel out, then the whole sum cancels."] -theorem prod_cancels_of_partition_cancels (R : Setoid α) [DecidableRel R.r] +theorem prod_cancels_of_partition_cancels (R : Setoid α) [DecidableRel R] (h : ∀ x ∈ s, ∏ a ∈ s.filter fun y => R y x, f a = 1) : ∏ x ∈ s, f x = 1 := by rw [prod_partition R, ← Finset.prod_eq_one] intro xbar xbar_in_s diff --git a/Mathlib/Algebra/CharZero/Quotient.lean b/Mathlib/Algebra/CharZero/Quotient.lean index 8a44865763c7e..79b566ed37ce1 100644 --- a/Mathlib/Algebra/CharZero/Quotient.lean +++ b/Mathlib/Algebra/CharZero/Quotient.lean @@ -51,11 +51,11 @@ namespace QuotientAddGroup theorem zmultiples_zsmul_eq_zsmul_iff {ψ θ : R ⧸ AddSubgroup.zmultiples p} {z : ℤ} (hz : z ≠ 0) : z • ψ = z • θ ↔ ∃ k : Fin z.natAbs, ψ = θ + ((k : ℕ) • (p / z) : R) := by - induction ψ using Quotient.inductionOn' - induction θ using Quotient.inductionOn' + induction ψ using Quotient.inductionOn + induction θ using Quotient.inductionOn -- Porting note: Introduced Zp notation to shorten lines let Zp := AddSubgroup.zmultiples p - have : (Quotient.mk'' : R → R ⧸ Zp) = ((↑) : R → R ⧸ Zp) := rfl + have : (Quotient.mk _ : R → R ⧸ Zp) = ((↑) : R → R ⧸ Zp) := rfl simp only [this] simp_rw [← QuotientAddGroup.mk_zsmul, ← QuotientAddGroup.mk_add, QuotientAddGroup.eq_iff_sub_mem, ← smul_sub, ← sub_sub] diff --git a/Mathlib/Data/List/Cycle.lean b/Mathlib/Data/List/Cycle.lean index 9c8cbaa143209..18607fde71b44 100644 --- a/Mathlib/Data/List/Cycle.lean +++ b/Mathlib/Data/List/Cycle.lean @@ -824,11 +824,11 @@ theorem chain_ne_nil (r : α → α → Prop) {l : List α} : theorem chain_map {β : Type*} {r : α → α → Prop} (f : β → α) {s : Cycle β} : Chain r (s.map f) ↔ Chain (fun a b => r (f a) (f b)) s := - Quotient.inductionOn' s fun l => by + Quotient.inductionOn s fun l => by cases' l with a l · rfl - dsimp only [Chain, ← mk''_eq_coe, Quotient.liftOn'_mk'', Cycle.map, Quotient.map', Quot.map, - Quotient.mk'', Quotient.liftOn', Quotient.liftOn, Quot.liftOn_mk, List.map] + dsimp only [Chain, Quotient.liftOn_mk, Cycle.map, Quotient.map', Quot.map, + Quotient.liftOn', Quotient.liftOn, Quot.liftOn_mk, List.map] rw [← concat_eq_append, ← List.map_concat, List.chain_map f] simp diff --git a/Mathlib/Data/Multiset/Basic.lean b/Mathlib/Data/Multiset/Basic.lean index e04ef7fa7a1dc..97366784110f6 100644 --- a/Mathlib/Data/Multiset/Basic.lean +++ b/Mathlib/Data/Multiset/Basic.lean @@ -67,6 +67,9 @@ theorem coe_eq_coe {l₁ l₂ : List α} : (l₁ : Multiset α) = l₂ ↔ l₁ instance [DecidableEq α] (l₁ l₂ : List α) : Decidable (l₁ ≈ l₂) := inferInstanceAs (Decidable (l₁ ~ l₂)) +instance [DecidableEq α] (l₁ l₂ : List α) : Decidable (isSetoid α l₁ l₂) := + inferInstanceAs (Decidable (l₁ ~ l₂)) + -- Porting note: `Quotient.recOnSubsingleton₂ s₁ s₂` was in parens which broke elaboration instance decidableEq [DecidableEq α] : DecidableEq (Multiset α) | s₁, s₂ => Quotient.recOnSubsingleton₂ s₁ s₂ fun _ _ => decidable_of_iff' _ Quotient.eq_iff_equiv diff --git a/Mathlib/Data/Quot.lean b/Mathlib/Data/Quot.lean index 2bec860ba95b9..8dfbafe4dbf99 100644 --- a/Mathlib/Data/Quot.lean +++ b/Mathlib/Data/Quot.lean @@ -546,8 +546,8 @@ several different quotient relations on a type, for example quotient groups, rin -- Porting note: Quotient.mk' is the equivalent of Lean 3's `Quotient.mk` /-- A version of `Quotient.mk` taking `{s : Setoid α}` as an implicit argument instead of an instance argument. -/ -protected def mk'' (a : α) : Quotient s₁ := - Quot.mk s₁.1 a +protected abbrev mk'' (a : α) : Quotient s₁ := + ⟦a⟧ /-- `Quotient.mk''` is a surjective function. -/ theorem surjective_Quotient_mk'' : Function.Surjective (Quotient.mk'' : α → Quotient s₁) := @@ -693,7 +693,6 @@ protected theorem eq' {s₁ : Setoid α} {a b : α} : @Quotient.mk' α s₁ a = @Quotient.mk' α s₁ b ↔ s₁ a b := Quotient.eq -@[simp] protected theorem eq'' {a b : α} : @Quotient.mk'' α s₁ a = Quotient.mk'' b ↔ s₁ a b := Quotient.eq @@ -725,7 +724,6 @@ protected theorem liftOn₂'_mk {t : Setoid β} (f : α → β → γ) (h) (a : Quotient.liftOn₂' (Quotient.mk s a) (Quotient.mk t b) f h = f a b := Quotient.liftOn₂'_mk'' _ _ _ _ -@[simp] theorem map'_mk {t : Setoid β} (f : α → β) (h) (x : α) : (Quotient.mk s x).map' f h = (Quotient.mk t (f x)) := rfl diff --git a/Mathlib/Data/Set/Image.lean b/Mathlib/Data/Set/Image.lean index 00387c99f2202..1da452c8964d3 100644 --- a/Mathlib/Data/Set/Image.lean +++ b/Mathlib/Data/Set/Image.lean @@ -836,7 +836,7 @@ theorem range_quotient_lift [s : Setoid ι] (hf) : theorem range_quotient_mk' {s : Setoid α} : range (Quotient.mk' : α → Quotient s) = univ := range_quot_mk _ -@[simp] lemma Quotient.range_mk'' {sa : Setoid α} : range (Quotient.mk'' (s₁ := sa)) = univ := +lemma Quotient.range_mk'' {sa : Setoid α} : range (Quotient.mk'' (s₁ := sa)) = univ := range_quotient_mk @[simp] diff --git a/Mathlib/GroupTheory/Coset/Basic.lean b/Mathlib/GroupTheory/Coset/Basic.lean index e7aa809a354ae..c5df256757ef5 100644 --- a/Mathlib/GroupTheory/Coset/Basic.lean +++ b/Mathlib/GroupTheory/Coset/Basic.lean @@ -491,7 +491,7 @@ open MulAction in lemma orbit_mk_eq_smul (x : α) : MulAction.orbitRel.Quotient.orbit (x : α ⧸ s) = x • s := by ext rw [orbitRel.Quotient.mem_orbit] - simpa [mem_smul_set_iff_inv_smul_mem, ← leftRel_apply] using Setoid.comm' _ + simpa [mem_smul_set_iff_inv_smul_mem, ← leftRel_apply, Quotient.eq''] using Setoid.comm' _ @[to_additive] lemma orbit_eq_out'_smul (x : α ⧸ s) : MulAction.orbitRel.Quotient.orbit x = x.out' • s := by @@ -530,7 +530,7 @@ noncomputable def groupEquivQuotientProdSubgroup : α ≃ (α ⧸ s) × s := show (_root_.Subtype fun x : α => Quotient.mk'' x = L) ≃ _root_.Subtype fun x : α => Quotient.mk'' x = Quotient.mk'' _ - simp [-Quotient.eq''] + simp rfl _ ≃ Σ _L : α ⧸ s, s := Equiv.sigmaCongrRight fun _ => leftCosetEquivSubgroup _ _ ≃ (α ⧸ s) × s := Equiv.sigmaEquivProd _ _ diff --git a/Mathlib/GroupTheory/GroupAction/Quotient.lean b/Mathlib/GroupTheory/GroupAction/Quotient.lean index f58416df124ac..e532945e65310 100644 --- a/Mathlib/GroupTheory/GroupAction/Quotient.lean +++ b/Mathlib/GroupTheory/GroupAction/Quotient.lean @@ -320,10 +320,10 @@ noncomputable def equivSubgroupOrbitsSetoidComap (H : Subgroup α) (ω : Ω) : toFun := fun q ↦ q.liftOn' (fun x ↦ ⟦⟨↑x, by simp only [Set.mem_preimage, Set.mem_singleton_iff] have hx := x.property - rwa [orbitRel.Quotient.mem_orbit, @Quotient.mk''_eq_mk] at hx⟩⟧) fun a b h ↦ by - simp only [← Quotient.eq'', Quotient.mk''_eq_mk, + rwa [orbitRel.Quotient.mem_orbit] at hx⟩⟧) fun a b h ↦ by + simp only [← Quotient.eq, orbitRel.Quotient.subgroup_quotient_eq_iff] at h - simp only [← Quotient.mk''_eq_mk, Quotient.eq''] at h ⊢ + simp only [Quotient.eq] at h ⊢ exact h invFun := fun q ↦ q.liftOn' (fun x ↦ ⟦⟨↑x, by have hx := x.property diff --git a/Mathlib/GroupTheory/PGroup.lean b/Mathlib/GroupTheory/PGroup.lean index d882d7e22ca40..eb9cdfbf182a9 100644 --- a/Mathlib/GroupTheory/PGroup.lean +++ b/Mathlib/GroupTheory/PGroup.lean @@ -179,7 +179,7 @@ theorem card_modEq_card_fixedPoints : Nat.card α ≡ Nat.card (fixedPoints G α rw [Nat.card_eq_fintype_card] at hk have : k = 0 := by contrapose! hb - simp [-Quotient.eq'', key, hk, hb] + simp [-Quotient.eq, key, hk, hb] exact ⟨⟨b, mem_fixedPoints_iff_card_orbit_eq_one.2 <| by rw [hk, this, pow_zero]⟩, Finset.mem_univ _, ne_of_eq_of_ne Nat.cast_one one_ne_zero, rfl⟩ From 64ac27ea68f1e00fc29e2ce69a1c73c348dcdfb4 Mon Sep 17 00:00:00 2001 From: Paul Lezeau Date: Sat, 19 Oct 2024 16:03:51 +0000 Subject: [PATCH 185/505] chore(Order/Bounds/Basic): Split long file into two shorter files (#16539) This PR transfers some of the content from `Order/Bounds/Basic` into a new file, `Order/Bounds/Image`. --- Mathlib.lean | 1 + Mathlib/Order/Bounds/Basic.lean | 482 --------------------------- Mathlib/Order/Bounds/Image.lean | 498 ++++++++++++++++++++++++++++ Mathlib/Order/Bounds/OrderIso.lean | 2 +- Mathlib/Order/CompleteLattice.lean | 1 + Mathlib/Order/GaloisConnection.lean | 2 +- 6 files changed, 502 insertions(+), 484 deletions(-) create mode 100644 Mathlib/Order/Bounds/Image.lean diff --git a/Mathlib.lean b/Mathlib.lean index 4c5d00a1fa732..5fcbb05388323 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3621,6 +3621,7 @@ import Mathlib.Order.Bounded import Mathlib.Order.BoundedOrder import Mathlib.Order.Bounds.Basic import Mathlib.Order.Bounds.Defs +import Mathlib.Order.Bounds.Image import Mathlib.Order.Bounds.OrderIso import Mathlib.Order.Category.BddDistLat import Mathlib.Order.Category.BddLat diff --git a/Mathlib/Order/Bounds/Basic.lean b/Mathlib/Order/Bounds/Basic.lean index fb6e31a954dff..66d0f19bb54b8 100644 --- a/Mathlib/Order/Bounds/Basic.lean +++ b/Mathlib/Order/Bounds/Basic.lean @@ -3,7 +3,6 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Yury Kudryashov -/ -import Mathlib.Data.Set.NAry import Mathlib.Order.Bounds.Defs import Mathlib.Order.Directed import Mathlib.Order.Interval.Set.Basic @@ -905,484 +904,3 @@ theorem IsGLB.exists_between' (h : IsGLB s a) (h' : a ∉ s) (hb : a < b) : ∃ ⟨c, hcs, hac.lt_of_ne fun hac => h' <| hac.symm ▸ hcs, hcb⟩ end LinearOrder - -/-! -### Images of upper/lower bounds under monotone functions --/ - - -namespace MonotoneOn - -variable [Preorder α] [Preorder β] {f : α → β} {s t : Set α} {a : α} - -theorem mem_upperBounds_image (Hf : MonotoneOn f t) (Hst : s ⊆ t) (Has : a ∈ upperBounds s) - (Hat : a ∈ t) : f a ∈ upperBounds (f '' s) := - forall_mem_image.2 fun _ H => Hf (Hst H) Hat (Has H) - -theorem mem_upperBounds_image_self (Hf : MonotoneOn f t) : - a ∈ upperBounds t → a ∈ t → f a ∈ upperBounds (f '' t) := - Hf.mem_upperBounds_image subset_rfl - -theorem mem_lowerBounds_image (Hf : MonotoneOn f t) (Hst : s ⊆ t) (Has : a ∈ lowerBounds s) - (Hat : a ∈ t) : f a ∈ lowerBounds (f '' s) := - forall_mem_image.2 fun _ H => Hf Hat (Hst H) (Has H) - -theorem mem_lowerBounds_image_self (Hf : MonotoneOn f t) : - a ∈ lowerBounds t → a ∈ t → f a ∈ lowerBounds (f '' t) := - Hf.mem_lowerBounds_image subset_rfl - -theorem image_upperBounds_subset_upperBounds_image (Hf : MonotoneOn f t) (Hst : s ⊆ t) : - f '' (upperBounds s ∩ t) ⊆ upperBounds (f '' s) := by - rintro _ ⟨a, ha, rfl⟩ - exact Hf.mem_upperBounds_image Hst ha.1 ha.2 - -theorem image_lowerBounds_subset_lowerBounds_image (Hf : MonotoneOn f t) (Hst : s ⊆ t) : - f '' (lowerBounds s ∩ t) ⊆ lowerBounds (f '' s) := - Hf.dual.image_upperBounds_subset_upperBounds_image Hst - -/-- The image under a monotone function on a set `t` of a subset which has an upper bound in `t` - is bounded above. -/ -theorem map_bddAbove (Hf : MonotoneOn f t) (Hst : s ⊆ t) : - (upperBounds s ∩ t).Nonempty → BddAbove (f '' s) := fun ⟨C, hs, ht⟩ => - ⟨f C, Hf.mem_upperBounds_image Hst hs ht⟩ - -/-- The image under a monotone function on a set `t` of a subset which has a lower bound in `t` - is bounded below. -/ -theorem map_bddBelow (Hf : MonotoneOn f t) (Hst : s ⊆ t) : - (lowerBounds s ∩ t).Nonempty → BddBelow (f '' s) := fun ⟨C, hs, ht⟩ => - ⟨f C, Hf.mem_lowerBounds_image Hst hs ht⟩ - -/-- A monotone map sends a least element of a set to a least element of its image. -/ -theorem map_isLeast (Hf : MonotoneOn f t) (Ha : IsLeast t a) : IsLeast (f '' t) (f a) := - ⟨mem_image_of_mem _ Ha.1, Hf.mem_lowerBounds_image_self Ha.2 Ha.1⟩ - -/-- A monotone map sends a greatest element of a set to a greatest element of its image. -/ -theorem map_isGreatest (Hf : MonotoneOn f t) (Ha : IsGreatest t a) : IsGreatest (f '' t) (f a) := - ⟨mem_image_of_mem _ Ha.1, Hf.mem_upperBounds_image_self Ha.2 Ha.1⟩ - -end MonotoneOn - -namespace AntitoneOn - -variable [Preorder α] [Preorder β] {f : α → β} {s t : Set α} {a : α} - -theorem mem_upperBounds_image (Hf : AntitoneOn f t) (Hst : s ⊆ t) (Has : a ∈ lowerBounds s) : - a ∈ t → f a ∈ upperBounds (f '' s) := - Hf.dual_right.mem_lowerBounds_image Hst Has - -theorem mem_upperBounds_image_self (Hf : AntitoneOn f t) : - a ∈ lowerBounds t → a ∈ t → f a ∈ upperBounds (f '' t) := - Hf.dual_right.mem_lowerBounds_image_self - -theorem mem_lowerBounds_image (Hf : AntitoneOn f t) (Hst : s ⊆ t) : - a ∈ upperBounds s → a ∈ t → f a ∈ lowerBounds (f '' s) := - Hf.dual_right.mem_upperBounds_image Hst - -theorem mem_lowerBounds_image_self (Hf : AntitoneOn f t) : - a ∈ upperBounds t → a ∈ t → f a ∈ lowerBounds (f '' t) := - Hf.dual_right.mem_upperBounds_image_self - -theorem image_lowerBounds_subset_upperBounds_image (Hf : AntitoneOn f t) (Hst : s ⊆ t) : - f '' (lowerBounds s ∩ t) ⊆ upperBounds (f '' s) := - Hf.dual_right.image_lowerBounds_subset_lowerBounds_image Hst - -theorem image_upperBounds_subset_lowerBounds_image (Hf : AntitoneOn f t) (Hst : s ⊆ t) : - f '' (upperBounds s ∩ t) ⊆ lowerBounds (f '' s) := - Hf.dual_right.image_upperBounds_subset_upperBounds_image Hst - -/-- The image under an antitone function of a set which is bounded above is bounded below. -/ -theorem map_bddAbove (Hf : AntitoneOn f t) (Hst : s ⊆ t) : - (upperBounds s ∩ t).Nonempty → BddBelow (f '' s) := - Hf.dual_right.map_bddAbove Hst - -/-- The image under an antitone function of a set which is bounded below is bounded above. -/ -theorem map_bddBelow (Hf : AntitoneOn f t) (Hst : s ⊆ t) : - (lowerBounds s ∩ t).Nonempty → BddAbove (f '' s) := - Hf.dual_right.map_bddBelow Hst - -/-- An antitone map sends a greatest element of a set to a least element of its image. -/ -theorem map_isGreatest (Hf : AntitoneOn f t) : IsGreatest t a → IsLeast (f '' t) (f a) := - Hf.dual_right.map_isGreatest - -/-- An antitone map sends a least element of a set to a greatest element of its image. -/ -theorem map_isLeast (Hf : AntitoneOn f t) : IsLeast t a → IsGreatest (f '' t) (f a) := - Hf.dual_right.map_isLeast - -end AntitoneOn - -namespace Monotone - -variable [Preorder α] [Preorder β] {f : α → β} (Hf : Monotone f) {a : α} {s : Set α} - -include Hf - -theorem mem_upperBounds_image (Ha : a ∈ upperBounds s) : f a ∈ upperBounds (f '' s) := - forall_mem_image.2 fun _ H => Hf (Ha H) - -theorem mem_lowerBounds_image (Ha : a ∈ lowerBounds s) : f a ∈ lowerBounds (f '' s) := - forall_mem_image.2 fun _ H => Hf (Ha H) - -theorem image_upperBounds_subset_upperBounds_image : - f '' upperBounds s ⊆ upperBounds (f '' s) := by - rintro _ ⟨a, ha, rfl⟩ - exact Hf.mem_upperBounds_image ha - -theorem image_lowerBounds_subset_lowerBounds_image : f '' lowerBounds s ⊆ lowerBounds (f '' s) := - Hf.dual.image_upperBounds_subset_upperBounds_image - -/-- The image under a monotone function of a set which is bounded above is bounded above. See also -`BddAbove.image2`. -/ -theorem map_bddAbove : BddAbove s → BddAbove (f '' s) - | ⟨C, hC⟩ => ⟨f C, Hf.mem_upperBounds_image hC⟩ - -/-- The image under a monotone function of a set which is bounded below is bounded below. See also -`BddBelow.image2`. -/ -theorem map_bddBelow : BddBelow s → BddBelow (f '' s) - | ⟨C, hC⟩ => ⟨f C, Hf.mem_lowerBounds_image hC⟩ - -/-- A monotone map sends a least element of a set to a least element of its image. -/ -theorem map_isLeast (Ha : IsLeast s a) : IsLeast (f '' s) (f a) := - ⟨mem_image_of_mem _ Ha.1, Hf.mem_lowerBounds_image Ha.2⟩ - -/-- A monotone map sends a greatest element of a set to a greatest element of its image. -/ -theorem map_isGreatest (Ha : IsGreatest s a) : IsGreatest (f '' s) (f a) := - ⟨mem_image_of_mem _ Ha.1, Hf.mem_upperBounds_image Ha.2⟩ - -end Monotone - -namespace Antitone - -variable [Preorder α] [Preorder β] {f : α → β} (hf : Antitone f) {a : α} {s : Set α} - -include hf - -theorem mem_upperBounds_image : a ∈ lowerBounds s → f a ∈ upperBounds (f '' s) := - hf.dual_right.mem_lowerBounds_image - -theorem mem_lowerBounds_image : a ∈ upperBounds s → f a ∈ lowerBounds (f '' s) := - hf.dual_right.mem_upperBounds_image - -theorem image_lowerBounds_subset_upperBounds_image : f '' lowerBounds s ⊆ upperBounds (f '' s) := - hf.dual_right.image_lowerBounds_subset_lowerBounds_image - -theorem image_upperBounds_subset_lowerBounds_image : f '' upperBounds s ⊆ lowerBounds (f '' s) := - hf.dual_right.image_upperBounds_subset_upperBounds_image - -/-- The image under an antitone function of a set which is bounded above is bounded below. -/ -theorem map_bddAbove : BddAbove s → BddBelow (f '' s) := - hf.dual_right.map_bddAbove - -/-- The image under an antitone function of a set which is bounded below is bounded above. -/ -theorem map_bddBelow : BddBelow s → BddAbove (f '' s) := - hf.dual_right.map_bddBelow - -/-- An antitone map sends a greatest element of a set to a least element of its image. -/ -theorem map_isGreatest : IsGreatest s a → IsLeast (f '' s) (f a) := - hf.dual_right.map_isGreatest - -/-- An antitone map sends a least element of a set to a greatest element of its image. -/ -theorem map_isLeast : IsLeast s a → IsGreatest (f '' s) (f a) := - hf.dual_right.map_isLeast - -end Antitone - -section Image2 - -variable [Preorder α] [Preorder β] [Preorder γ] {f : α → β → γ} {s : Set α} {t : Set β} {a : α} - {b : β} - -section MonotoneMonotone - -variable (h₀ : ∀ b, Monotone (swap f b)) (h₁ : ∀ a, Monotone (f a)) - -include h₀ h₁ - -theorem mem_upperBounds_image2 (ha : a ∈ upperBounds s) (hb : b ∈ upperBounds t) : - f a b ∈ upperBounds (image2 f s t) := - forall_image2_iff.2 fun _ hx _ hy => (h₀ _ <| ha hx).trans <| h₁ _ <| hb hy - -theorem mem_lowerBounds_image2 (ha : a ∈ lowerBounds s) (hb : b ∈ lowerBounds t) : - f a b ∈ lowerBounds (image2 f s t) := - forall_image2_iff.2 fun _ hx _ hy => (h₀ _ <| ha hx).trans <| h₁ _ <| hb hy - -theorem image2_upperBounds_upperBounds_subset : - image2 f (upperBounds s) (upperBounds t) ⊆ upperBounds (image2 f s t) := - image2_subset_iff.2 fun _ ha _ hb ↦ mem_upperBounds_image2 h₀ h₁ ha hb - -theorem image2_lowerBounds_lowerBounds_subset : - image2 f (lowerBounds s) (lowerBounds t) ⊆ lowerBounds (image2 f s t) := - image2_subset_iff.2 fun _ ha _ hb ↦ mem_lowerBounds_image2 h₀ h₁ ha hb - -/-- See also `Monotone.map_bddAbove`. -/ -protected theorem BddAbove.image2 : - BddAbove s → BddAbove t → BddAbove (image2 f s t) := by - rintro ⟨a, ha⟩ ⟨b, hb⟩ - exact ⟨f a b, mem_upperBounds_image2 h₀ h₁ ha hb⟩ - -/-- See also `Monotone.map_bddBelow`. -/ -protected theorem BddBelow.image2 : - BddBelow s → BddBelow t → BddBelow (image2 f s t) := by - rintro ⟨a, ha⟩ ⟨b, hb⟩ - exact ⟨f a b, mem_lowerBounds_image2 h₀ h₁ ha hb⟩ - -protected theorem IsGreatest.image2 (ha : IsGreatest s a) (hb : IsGreatest t b) : - IsGreatest (image2 f s t) (f a b) := - ⟨mem_image2_of_mem ha.1 hb.1, mem_upperBounds_image2 h₀ h₁ ha.2 hb.2⟩ - -protected theorem IsLeast.image2 (ha : IsLeast s a) (hb : IsLeast t b) : - IsLeast (image2 f s t) (f a b) := - ⟨mem_image2_of_mem ha.1 hb.1, mem_lowerBounds_image2 h₀ h₁ ha.2 hb.2⟩ - -end MonotoneMonotone - -section MonotoneAntitone - -variable (h₀ : ∀ b, Monotone (swap f b)) (h₁ : ∀ a, Antitone (f a)) - -include h₀ h₁ - -theorem mem_upperBounds_image2_of_mem_upperBounds_of_mem_lowerBounds (ha : a ∈ upperBounds s) - (hb : b ∈ lowerBounds t) : f a b ∈ upperBounds (image2 f s t) := - forall_image2_iff.2 fun _ hx _ hy => (h₀ _ <| ha hx).trans <| h₁ _ <| hb hy - -theorem mem_lowerBounds_image2_of_mem_lowerBounds_of_mem_upperBounds (ha : a ∈ lowerBounds s) - (hb : b ∈ upperBounds t) : f a b ∈ lowerBounds (image2 f s t) := - forall_image2_iff.2 fun _ hx _ hy => (h₀ _ <| ha hx).trans <| h₁ _ <| hb hy - -theorem image2_upperBounds_lowerBounds_subset_upperBounds_image2 : - image2 f (upperBounds s) (lowerBounds t) ⊆ upperBounds (image2 f s t) := - image2_subset_iff.2 fun _ ha _ hb ↦ - mem_upperBounds_image2_of_mem_upperBounds_of_mem_lowerBounds h₀ h₁ ha hb - -theorem image2_lowerBounds_upperBounds_subset_lowerBounds_image2 : - image2 f (lowerBounds s) (upperBounds t) ⊆ lowerBounds (image2 f s t) := - image2_subset_iff.2 fun _ ha _ hb ↦ - mem_lowerBounds_image2_of_mem_lowerBounds_of_mem_upperBounds h₀ h₁ ha hb - -theorem BddAbove.bddAbove_image2_of_bddBelow : - BddAbove s → BddBelow t → BddAbove (Set.image2 f s t) := by - rintro ⟨a, ha⟩ ⟨b, hb⟩ - exact ⟨f a b, mem_upperBounds_image2_of_mem_upperBounds_of_mem_lowerBounds h₀ h₁ ha hb⟩ - -theorem BddBelow.bddBelow_image2_of_bddAbove : - BddBelow s → BddAbove t → BddBelow (Set.image2 f s t) := by - rintro ⟨a, ha⟩ ⟨b, hb⟩ - exact ⟨f a b, mem_lowerBounds_image2_of_mem_lowerBounds_of_mem_upperBounds h₀ h₁ ha hb⟩ - -theorem IsGreatest.isGreatest_image2_of_isLeast (ha : IsGreatest s a) (hb : IsLeast t b) : - IsGreatest (Set.image2 f s t) (f a b) := - ⟨mem_image2_of_mem ha.1 hb.1, - mem_upperBounds_image2_of_mem_upperBounds_of_mem_lowerBounds h₀ h₁ ha.2 hb.2⟩ - -theorem IsLeast.isLeast_image2_of_isGreatest (ha : IsLeast s a) (hb : IsGreatest t b) : - IsLeast (Set.image2 f s t) (f a b) := - ⟨mem_image2_of_mem ha.1 hb.1, - mem_lowerBounds_image2_of_mem_lowerBounds_of_mem_upperBounds h₀ h₁ ha.2 hb.2⟩ - -end MonotoneAntitone - -section AntitoneAntitone - -variable (h₀ : ∀ b, Antitone (swap f b)) (h₁ : ∀ a, Antitone (f a)) - -include h₀ h₁ - -theorem mem_upperBounds_image2_of_mem_lowerBounds (ha : a ∈ lowerBounds s) - (hb : b ∈ lowerBounds t) : f a b ∈ upperBounds (image2 f s t) := - forall_image2_iff.2 fun _ hx _ hy => (h₀ _ <| ha hx).trans <| h₁ _ <| hb hy - -theorem mem_lowerBounds_image2_of_mem_upperBounds (ha : a ∈ upperBounds s) - (hb : b ∈ upperBounds t) : f a b ∈ lowerBounds (image2 f s t) := - forall_image2_iff.2 fun _ hx _ hy => (h₀ _ <| ha hx).trans <| h₁ _ <| hb hy - -theorem image2_upperBounds_upperBounds_subset_upperBounds_image2 : - image2 f (lowerBounds s) (lowerBounds t) ⊆ upperBounds (image2 f s t) := - image2_subset_iff.2 fun _ ha _ hb ↦ - mem_upperBounds_image2_of_mem_lowerBounds h₀ h₁ ha hb - -theorem image2_lowerBounds_lowerBounds_subset_lowerBounds_image2 : - image2 f (upperBounds s) (upperBounds t) ⊆ lowerBounds (image2 f s t) := - image2_subset_iff.2 fun _ ha _ hb ↦ - mem_lowerBounds_image2_of_mem_upperBounds h₀ h₁ ha hb - -theorem BddBelow.image2_bddAbove : BddBelow s → BddBelow t → BddAbove (Set.image2 f s t) := by - rintro ⟨a, ha⟩ ⟨b, hb⟩ - exact ⟨f a b, mem_upperBounds_image2_of_mem_lowerBounds h₀ h₁ ha hb⟩ - -theorem BddAbove.image2_bddBelow : BddAbove s → BddAbove t → BddBelow (Set.image2 f s t) := by - rintro ⟨a, ha⟩ ⟨b, hb⟩ - exact ⟨f a b, mem_lowerBounds_image2_of_mem_upperBounds h₀ h₁ ha hb⟩ - -theorem IsLeast.isGreatest_image2 (ha : IsLeast s a) (hb : IsLeast t b) : - IsGreatest (Set.image2 f s t) (f a b) := - ⟨mem_image2_of_mem ha.1 hb.1, mem_upperBounds_image2_of_mem_lowerBounds h₀ h₁ ha.2 hb.2⟩ - -theorem IsGreatest.isLeast_image2 (ha : IsGreatest s a) (hb : IsGreatest t b) : - IsLeast (Set.image2 f s t) (f a b) := - ⟨mem_image2_of_mem ha.1 hb.1, mem_lowerBounds_image2_of_mem_upperBounds h₀ h₁ ha.2 hb.2⟩ - -end AntitoneAntitone - -section AntitoneMonotone - -variable (h₀ : ∀ b, Antitone (swap f b)) (h₁ : ∀ a, Monotone (f a)) - -include h₀ h₁ - -theorem mem_upperBounds_image2_of_mem_upperBounds_of_mem_upperBounds (ha : a ∈ lowerBounds s) - (hb : b ∈ upperBounds t) : f a b ∈ upperBounds (image2 f s t) := - forall_image2_iff.2 fun _ hx _ hy => (h₀ _ <| ha hx).trans <| h₁ _ <| hb hy - -theorem mem_lowerBounds_image2_of_mem_lowerBounds_of_mem_lowerBounds (ha : a ∈ upperBounds s) - (hb : b ∈ lowerBounds t) : f a b ∈ lowerBounds (image2 f s t) := - forall_image2_iff.2 fun _ hx _ hy => (h₀ _ <| ha hx).trans <| h₁ _ <| hb hy - -theorem image2_lowerBounds_upperBounds_subset_upperBounds_image2 : - image2 f (lowerBounds s) (upperBounds t) ⊆ upperBounds (image2 f s t) := - image2_subset_iff.2 fun _ ha _ hb ↦ - mem_upperBounds_image2_of_mem_upperBounds_of_mem_upperBounds h₀ h₁ ha hb - -theorem image2_upperBounds_lowerBounds_subset_lowerBounds_image2 : - image2 f (upperBounds s) (lowerBounds t) ⊆ lowerBounds (image2 f s t) := - image2_subset_iff.2 fun _ ha _ hb ↦ - mem_lowerBounds_image2_of_mem_lowerBounds_of_mem_lowerBounds h₀ h₁ ha hb - -theorem BddBelow.bddAbove_image2_of_bddAbove : - BddBelow s → BddAbove t → BddAbove (Set.image2 f s t) := by - rintro ⟨a, ha⟩ ⟨b, hb⟩ - exact ⟨f a b, mem_upperBounds_image2_of_mem_upperBounds_of_mem_upperBounds h₀ h₁ ha hb⟩ - -theorem BddAbove.bddBelow_image2_of_bddAbove : - BddAbove s → BddBelow t → BddBelow (Set.image2 f s t) := by - rintro ⟨a, ha⟩ ⟨b, hb⟩ - exact ⟨f a b, mem_lowerBounds_image2_of_mem_lowerBounds_of_mem_lowerBounds h₀ h₁ ha hb⟩ - -theorem IsLeast.isGreatest_image2_of_isGreatest (ha : IsLeast s a) (hb : IsGreatest t b) : - IsGreatest (Set.image2 f s t) (f a b) := - ⟨mem_image2_of_mem ha.1 hb.1, - mem_upperBounds_image2_of_mem_upperBounds_of_mem_upperBounds h₀ h₁ ha.2 hb.2⟩ - -theorem IsGreatest.isLeast_image2_of_isLeast (ha : IsGreatest s a) (hb : IsLeast t b) : - IsLeast (Set.image2 f s t) (f a b) := - ⟨mem_image2_of_mem ha.1 hb.1, - mem_lowerBounds_image2_of_mem_lowerBounds_of_mem_lowerBounds h₀ h₁ ha.2 hb.2⟩ - -end AntitoneMonotone - -end Image2 - -section Prod - -variable {α β : Type*} [Preorder α] [Preorder β] - -lemma bddAbove_prod {s : Set (α × β)} : - BddAbove s ↔ BddAbove (Prod.fst '' s) ∧ BddAbove (Prod.snd '' s) := - ⟨fun ⟨p, hp⟩ ↦ ⟨⟨p.1, forall_mem_image.2 fun _q hq ↦ (hp hq).1⟩, - ⟨p.2, forall_mem_image.2 fun _q hq ↦ (hp hq).2⟩⟩, - fun ⟨⟨x, hx⟩, ⟨y, hy⟩⟩ ↦ ⟨⟨x, y⟩, fun _p hp ↦ - ⟨hx <| mem_image_of_mem _ hp, hy <| mem_image_of_mem _ hp⟩⟩⟩ - -lemma bddBelow_prod {s : Set (α × β)} : - BddBelow s ↔ BddBelow (Prod.fst '' s) ∧ BddBelow (Prod.snd '' s) := - bddAbove_prod (α := αᵒᵈ) (β := βᵒᵈ) - -lemma bddAbove_range_prod {F : ι → α × β} : - BddAbove (range F) ↔ BddAbove (range <| Prod.fst ∘ F) ∧ BddAbove (range <| Prod.snd ∘ F) := by - simp only [bddAbove_prod, ← range_comp] - -lemma bddBelow_range_prod {F : ι → α × β} : - BddBelow (range F) ↔ BddBelow (range <| Prod.fst ∘ F) ∧ BddBelow (range <| Prod.snd ∘ F) := - bddAbove_range_prod (α := αᵒᵈ) (β := βᵒᵈ) - -theorem isLUB_prod {s : Set (α × β)} (p : α × β) : - IsLUB s p ↔ IsLUB (Prod.fst '' s) p.1 ∧ IsLUB (Prod.snd '' s) p.2 := by - refine - ⟨fun H => - ⟨⟨monotone_fst.mem_upperBounds_image H.1, fun a ha => ?_⟩, - ⟨monotone_snd.mem_upperBounds_image H.1, fun a ha => ?_⟩⟩, - fun H => ⟨?_, ?_⟩⟩ - · suffices h : (a, p.2) ∈ upperBounds s from (H.2 h).1 - exact fun q hq => ⟨ha <| mem_image_of_mem _ hq, (H.1 hq).2⟩ - · suffices h : (p.1, a) ∈ upperBounds s from (H.2 h).2 - exact fun q hq => ⟨(H.1 hq).1, ha <| mem_image_of_mem _ hq⟩ - · exact fun q hq => ⟨H.1.1 <| mem_image_of_mem _ hq, H.2.1 <| mem_image_of_mem _ hq⟩ - · exact fun q hq => - ⟨H.1.2 <| monotone_fst.mem_upperBounds_image hq, - H.2.2 <| monotone_snd.mem_upperBounds_image hq⟩ - -theorem isGLB_prod {s : Set (α × β)} (p : α × β) : - IsGLB s p ↔ IsGLB (Prod.fst '' s) p.1 ∧ IsGLB (Prod.snd '' s) p.2 := - @isLUB_prod αᵒᵈ βᵒᵈ _ _ _ _ - -end Prod - - -section Pi - -variable {π : α → Type*} [∀ a, Preorder (π a)] - -lemma bddAbove_pi {s : Set (∀ a, π a)} : - BddAbove s ↔ ∀ a, BddAbove (Function.eval a '' s) := - ⟨fun ⟨f, hf⟩ a ↦ ⟨f a, forall_mem_image.2 fun _ hg ↦ hf hg a⟩, - fun h ↦ ⟨fun a ↦ (h a).some, fun _ hg a ↦ (h a).some_mem <| mem_image_of_mem _ hg⟩⟩ - -lemma bddBelow_pi {s : Set (∀ a, π a)} : - BddBelow s ↔ ∀ a, BddBelow (Function.eval a '' s) := - bddAbove_pi (π := fun a ↦ (π a)ᵒᵈ) - -lemma bddAbove_range_pi {F : ι → ∀ a, π a} : - BddAbove (range F) ↔ ∀ a, BddAbove (range fun i ↦ F i a) := by - simp only [bddAbove_pi, ← range_comp] - rfl - -lemma bddBelow_range_pi {F : ι → ∀ a, π a} : - BddBelow (range F) ↔ ∀ a, BddBelow (range fun i ↦ F i a) := - bddAbove_range_pi (π := fun a ↦ (π a)ᵒᵈ) - -theorem isLUB_pi {s : Set (∀ a, π a)} {f : ∀ a, π a} : - IsLUB s f ↔ ∀ a, IsLUB (Function.eval a '' s) (f a) := by - classical - refine - ⟨fun H a => ⟨(Function.monotone_eval a).mem_upperBounds_image H.1, fun b hb => ?_⟩, fun H => - ⟨?_, ?_⟩⟩ - · suffices h : Function.update f a b ∈ upperBounds s from Function.update_same a b f ▸ H.2 h a - exact fun g hg => le_update_iff.2 ⟨hb <| mem_image_of_mem _ hg, fun i _ => H.1 hg i⟩ - · exact fun g hg a => (H a).1 (mem_image_of_mem _ hg) - · exact fun g hg a => (H a).2 ((Function.monotone_eval a).mem_upperBounds_image hg) - -theorem isGLB_pi {s : Set (∀ a, π a)} {f : ∀ a, π a} : - IsGLB s f ↔ ∀ a, IsGLB (Function.eval a '' s) (f a) := - @isLUB_pi α (fun a => (π a)ᵒᵈ) _ s f - -end Pi - -theorem IsGLB.of_image [Preorder α] [Preorder β] {f : α → β} (hf : ∀ {x y}, f x ≤ f y ↔ x ≤ y) - {s : Set α} {x : α} (hx : IsGLB (f '' s) (f x)) : IsGLB s x := - ⟨fun _ hy => hf.1 <| hx.1 <| mem_image_of_mem _ hy, fun _ hy => - hf.1 <| hx.2 <| Monotone.mem_lowerBounds_image (fun _ _ => hf.2) hy⟩ - -theorem IsLUB.of_image [Preorder α] [Preorder β] {f : α → β} (hf : ∀ {x y}, f x ≤ f y ↔ x ≤ y) - {s : Set α} {x : α} (hx : IsLUB (f '' s) (f x)) : IsLUB s x := - ⟨fun _ hy => hf.1 <| hx.1 <| mem_image_of_mem _ hy, fun _ hy => - hf.1 <| hx.2 <| Monotone.mem_upperBounds_image (fun _ _ => hf.2) hy⟩ - -lemma BddAbove.range_mono [Preorder β] {f : α → β} (g : α → β) (h : ∀ a, f a ≤ g a) - (hbdd : BddAbove (range g)) : BddAbove (range f) := by - obtain ⟨C, hC⟩ := hbdd - use C - rintro - ⟨x, rfl⟩ - exact (h x).trans (hC <| mem_range_self x) - -lemma BddBelow.range_mono [Preorder β] (f : α → β) {g : α → β} (h : ∀ a, f a ≤ g a) - (hbdd : BddBelow (range f)) : BddBelow (range g) := - BddAbove.range_mono (β := βᵒᵈ) f h hbdd - -lemma BddAbove.range_comp {γ : Type*} [Preorder β] [Preorder γ] {f : α → β} {g : β → γ} - (hf : BddAbove (range f)) (hg : Monotone g) : BddAbove (range (fun x => g (f x))) := by - change BddAbove (range (g ∘ f)) - simpa only [Set.range_comp] using hg.map_bddAbove hf - -lemma BddBelow.range_comp {γ : Type*} [Preorder β] [Preorder γ] {f : α → β} {g : β → γ} - (hf : BddBelow (range f)) (hg : Monotone g) : BddBelow (range (fun x => g (f x))) := by - change BddBelow (range (g ∘ f)) - simpa only [Set.range_comp] using hg.map_bddBelow hf diff --git a/Mathlib/Order/Bounds/Image.lean b/Mathlib/Order/Bounds/Image.lean new file mode 100644 index 0000000000000..910c466332520 --- /dev/null +++ b/Mathlib/Order/Bounds/Image.lean @@ -0,0 +1,498 @@ +/- +Copyright (c) 2017 Paul Lezeau. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johannes Hölzl, Yury Kudryashov, Paul Lezeau +-/ +import Mathlib.Data.Set.NAry +import Mathlib.Order.Bounds.Defs + +/-! + +# Images of upper/lower bounds under monotone functions + +In this file we prove various results about the behaviour of bounds under monotone/antitone maps. +-/ + +open Function Set + +open OrderDual (toDual ofDual) + +universe u v w x + +variable {α : Type u} {β : Type v} {γ : Type w} {ι : Sort x} + +namespace MonotoneOn + +variable [Preorder α] [Preorder β] {f : α → β} {s t : Set α} {a : α} + +theorem mem_upperBounds_image (Hf : MonotoneOn f t) (Hst : s ⊆ t) (Has : a ∈ upperBounds s) + (Hat : a ∈ t) : f a ∈ upperBounds (f '' s) := + forall_mem_image.2 fun _ H => Hf (Hst H) Hat (Has H) + +theorem mem_upperBounds_image_self (Hf : MonotoneOn f t) : + a ∈ upperBounds t → a ∈ t → f a ∈ upperBounds (f '' t) := + Hf.mem_upperBounds_image subset_rfl + +theorem mem_lowerBounds_image (Hf : MonotoneOn f t) (Hst : s ⊆ t) (Has : a ∈ lowerBounds s) + (Hat : a ∈ t) : f a ∈ lowerBounds (f '' s) := + forall_mem_image.2 fun _ H => Hf Hat (Hst H) (Has H) + +theorem mem_lowerBounds_image_self (Hf : MonotoneOn f t) : + a ∈ lowerBounds t → a ∈ t → f a ∈ lowerBounds (f '' t) := + Hf.mem_lowerBounds_image subset_rfl + +theorem image_upperBounds_subset_upperBounds_image (Hf : MonotoneOn f t) (Hst : s ⊆ t) : + f '' (upperBounds s ∩ t) ⊆ upperBounds (f '' s) := by + rintro _ ⟨a, ha, rfl⟩ + exact Hf.mem_upperBounds_image Hst ha.1 ha.2 + +theorem image_lowerBounds_subset_lowerBounds_image (Hf : MonotoneOn f t) (Hst : s ⊆ t) : + f '' (lowerBounds s ∩ t) ⊆ lowerBounds (f '' s) := + Hf.dual.image_upperBounds_subset_upperBounds_image Hst + +/-- The image under a monotone function on a set `t` of a subset which has an upper bound in `t` + is bounded above. -/ +theorem map_bddAbove (Hf : MonotoneOn f t) (Hst : s ⊆ t) : + (upperBounds s ∩ t).Nonempty → BddAbove (f '' s) := fun ⟨C, hs, ht⟩ => + ⟨f C, Hf.mem_upperBounds_image Hst hs ht⟩ + +/-- The image under a monotone function on a set `t` of a subset which has a lower bound in `t` + is bounded below. -/ +theorem map_bddBelow (Hf : MonotoneOn f t) (Hst : s ⊆ t) : + (lowerBounds s ∩ t).Nonempty → BddBelow (f '' s) := fun ⟨C, hs, ht⟩ => + ⟨f C, Hf.mem_lowerBounds_image Hst hs ht⟩ + +/-- A monotone map sends a least element of a set to a least element of its image. -/ +theorem map_isLeast (Hf : MonotoneOn f t) (Ha : IsLeast t a) : IsLeast (f '' t) (f a) := + ⟨mem_image_of_mem _ Ha.1, Hf.mem_lowerBounds_image_self Ha.2 Ha.1⟩ + +/-- A monotone map sends a greatest element of a set to a greatest element of its image. -/ +theorem map_isGreatest (Hf : MonotoneOn f t) (Ha : IsGreatest t a) : IsGreatest (f '' t) (f a) := + ⟨mem_image_of_mem _ Ha.1, Hf.mem_upperBounds_image_self Ha.2 Ha.1⟩ + +end MonotoneOn + +namespace AntitoneOn + +variable [Preorder α] [Preorder β] {f : α → β} {s t : Set α} {a : α} + +theorem mem_upperBounds_image (Hf : AntitoneOn f t) (Hst : s ⊆ t) (Has : a ∈ lowerBounds s) : + a ∈ t → f a ∈ upperBounds (f '' s) := + Hf.dual_right.mem_lowerBounds_image Hst Has + +theorem mem_upperBounds_image_self (Hf : AntitoneOn f t) : + a ∈ lowerBounds t → a ∈ t → f a ∈ upperBounds (f '' t) := + Hf.dual_right.mem_lowerBounds_image_self + +theorem mem_lowerBounds_image (Hf : AntitoneOn f t) (Hst : s ⊆ t) : + a ∈ upperBounds s → a ∈ t → f a ∈ lowerBounds (f '' s) := + Hf.dual_right.mem_upperBounds_image Hst + +theorem mem_lowerBounds_image_self (Hf : AntitoneOn f t) : + a ∈ upperBounds t → a ∈ t → f a ∈ lowerBounds (f '' t) := + Hf.dual_right.mem_upperBounds_image_self + +theorem image_lowerBounds_subset_upperBounds_image (Hf : AntitoneOn f t) (Hst : s ⊆ t) : + f '' (lowerBounds s ∩ t) ⊆ upperBounds (f '' s) := + Hf.dual_right.image_lowerBounds_subset_lowerBounds_image Hst + +theorem image_upperBounds_subset_lowerBounds_image (Hf : AntitoneOn f t) (Hst : s ⊆ t) : + f '' (upperBounds s ∩ t) ⊆ lowerBounds (f '' s) := + Hf.dual_right.image_upperBounds_subset_upperBounds_image Hst + +/-- The image under an antitone function of a set which is bounded above is bounded below. -/ +theorem map_bddAbove (Hf : AntitoneOn f t) (Hst : s ⊆ t) : + (upperBounds s ∩ t).Nonempty → BddBelow (f '' s) := + Hf.dual_right.map_bddAbove Hst + +/-- The image under an antitone function of a set which is bounded below is bounded above. -/ +theorem map_bddBelow (Hf : AntitoneOn f t) (Hst : s ⊆ t) : + (lowerBounds s ∩ t).Nonempty → BddAbove (f '' s) := + Hf.dual_right.map_bddBelow Hst + +/-- An antitone map sends a greatest element of a set to a least element of its image. -/ +theorem map_isGreatest (Hf : AntitoneOn f t) : IsGreatest t a → IsLeast (f '' t) (f a) := + Hf.dual_right.map_isGreatest + +/-- An antitone map sends a least element of a set to a greatest element of its image. -/ +theorem map_isLeast (Hf : AntitoneOn f t) : IsLeast t a → IsGreatest (f '' t) (f a) := + Hf.dual_right.map_isLeast + +end AntitoneOn + +namespace Monotone + +variable [Preorder α] [Preorder β] {f : α → β} (Hf : Monotone f) {a : α} {s : Set α} + +include Hf + +theorem mem_upperBounds_image (Ha : a ∈ upperBounds s) : f a ∈ upperBounds (f '' s) := + forall_mem_image.2 fun _ H => Hf (Ha H) + +theorem mem_lowerBounds_image (Ha : a ∈ lowerBounds s) : f a ∈ lowerBounds (f '' s) := + forall_mem_image.2 fun _ H => Hf (Ha H) + +theorem image_upperBounds_subset_upperBounds_image : + f '' upperBounds s ⊆ upperBounds (f '' s) := by + rintro _ ⟨a, ha, rfl⟩ + exact Hf.mem_upperBounds_image ha + +theorem image_lowerBounds_subset_lowerBounds_image : f '' lowerBounds s ⊆ lowerBounds (f '' s) := + Hf.dual.image_upperBounds_subset_upperBounds_image + +/-- The image under a monotone function of a set which is bounded above is bounded above. See also +`BddAbove.image2`. -/ +theorem map_bddAbove : BddAbove s → BddAbove (f '' s) + | ⟨C, hC⟩ => ⟨f C, Hf.mem_upperBounds_image hC⟩ + +/-- The image under a monotone function of a set which is bounded below is bounded below. See also +`BddBelow.image2`. -/ +theorem map_bddBelow : BddBelow s → BddBelow (f '' s) + | ⟨C, hC⟩ => ⟨f C, Hf.mem_lowerBounds_image hC⟩ + +/-- A monotone map sends a least element of a set to a least element of its image. -/ +theorem map_isLeast (Ha : IsLeast s a) : IsLeast (f '' s) (f a) := + ⟨mem_image_of_mem _ Ha.1, Hf.mem_lowerBounds_image Ha.2⟩ + +/-- A monotone map sends a greatest element of a set to a greatest element of its image. -/ +theorem map_isGreatest (Ha : IsGreatest s a) : IsGreatest (f '' s) (f a) := + ⟨mem_image_of_mem _ Ha.1, Hf.mem_upperBounds_image Ha.2⟩ + +end Monotone + +namespace Antitone + +variable [Preorder α] [Preorder β] {f : α → β} (hf : Antitone f) {a : α} {s : Set α} + +include hf + +theorem mem_upperBounds_image : a ∈ lowerBounds s → f a ∈ upperBounds (f '' s) := + hf.dual_right.mem_lowerBounds_image + +theorem mem_lowerBounds_image : a ∈ upperBounds s → f a ∈ lowerBounds (f '' s) := + hf.dual_right.mem_upperBounds_image + +theorem image_lowerBounds_subset_upperBounds_image : f '' lowerBounds s ⊆ upperBounds (f '' s) := + hf.dual_right.image_lowerBounds_subset_lowerBounds_image + +theorem image_upperBounds_subset_lowerBounds_image : f '' upperBounds s ⊆ lowerBounds (f '' s) := + hf.dual_right.image_upperBounds_subset_upperBounds_image + +/-- The image under an antitone function of a set which is bounded above is bounded below. -/ +theorem map_bddAbove : BddAbove s → BddBelow (f '' s) := + hf.dual_right.map_bddAbove + +/-- The image under an antitone function of a set which is bounded below is bounded above. -/ +theorem map_bddBelow : BddBelow s → BddAbove (f '' s) := + hf.dual_right.map_bddBelow + +/-- An antitone map sends a greatest element of a set to a least element of its image. -/ +theorem map_isGreatest : IsGreatest s a → IsLeast (f '' s) (f a) := + hf.dual_right.map_isGreatest + +/-- An antitone map sends a least element of a set to a greatest element of its image. -/ +theorem map_isLeast : IsLeast s a → IsGreatest (f '' s) (f a) := + hf.dual_right.map_isLeast + +end Antitone + +section Image2 + +variable [Preorder α] [Preorder β] [Preorder γ] {f : α → β → γ} {s : Set α} {t : Set β} {a : α} + {b : β} + +section MonotoneMonotone + +variable (h₀ : ∀ b, Monotone (swap f b)) (h₁ : ∀ a, Monotone (f a)) + +include h₀ h₁ + +theorem mem_upperBounds_image2 (ha : a ∈ upperBounds s) (hb : b ∈ upperBounds t) : + f a b ∈ upperBounds (image2 f s t) := + forall_image2_iff.2 fun _ hx _ hy => (h₀ _ <| ha hx).trans <| h₁ _ <| hb hy + +theorem mem_lowerBounds_image2 (ha : a ∈ lowerBounds s) (hb : b ∈ lowerBounds t) : + f a b ∈ lowerBounds (image2 f s t) := + forall_image2_iff.2 fun _ hx _ hy => (h₀ _ <| ha hx).trans <| h₁ _ <| hb hy + +theorem image2_upperBounds_upperBounds_subset : + image2 f (upperBounds s) (upperBounds t) ⊆ upperBounds (image2 f s t) := + image2_subset_iff.2 fun _ ha _ hb ↦ mem_upperBounds_image2 h₀ h₁ ha hb + +theorem image2_lowerBounds_lowerBounds_subset : + image2 f (lowerBounds s) (lowerBounds t) ⊆ lowerBounds (image2 f s t) := + image2_subset_iff.2 fun _ ha _ hb ↦ mem_lowerBounds_image2 h₀ h₁ ha hb + +/-- See also `Monotone.map_bddAbove`. -/ +protected theorem BddAbove.image2 : + BddAbove s → BddAbove t → BddAbove (image2 f s t) := by + rintro ⟨a, ha⟩ ⟨b, hb⟩ + exact ⟨f a b, mem_upperBounds_image2 h₀ h₁ ha hb⟩ + +/-- See also `Monotone.map_bddBelow`. -/ +protected theorem BddBelow.image2 : + BddBelow s → BddBelow t → BddBelow (image2 f s t) := by + rintro ⟨a, ha⟩ ⟨b, hb⟩ + exact ⟨f a b, mem_lowerBounds_image2 h₀ h₁ ha hb⟩ + +protected theorem IsGreatest.image2 (ha : IsGreatest s a) (hb : IsGreatest t b) : + IsGreatest (image2 f s t) (f a b) := + ⟨mem_image2_of_mem ha.1 hb.1, mem_upperBounds_image2 h₀ h₁ ha.2 hb.2⟩ + +protected theorem IsLeast.image2 (ha : IsLeast s a) (hb : IsLeast t b) : + IsLeast (image2 f s t) (f a b) := + ⟨mem_image2_of_mem ha.1 hb.1, mem_lowerBounds_image2 h₀ h₁ ha.2 hb.2⟩ + +end MonotoneMonotone + +section MonotoneAntitone + +variable (h₀ : ∀ b, Monotone (swap f b)) (h₁ : ∀ a, Antitone (f a)) + +include h₀ h₁ + +theorem mem_upperBounds_image2_of_mem_upperBounds_of_mem_lowerBounds (ha : a ∈ upperBounds s) + (hb : b ∈ lowerBounds t) : f a b ∈ upperBounds (image2 f s t) := + forall_image2_iff.2 fun _ hx _ hy => (h₀ _ <| ha hx).trans <| h₁ _ <| hb hy + +theorem mem_lowerBounds_image2_of_mem_lowerBounds_of_mem_upperBounds (ha : a ∈ lowerBounds s) + (hb : b ∈ upperBounds t) : f a b ∈ lowerBounds (image2 f s t) := + forall_image2_iff.2 fun _ hx _ hy => (h₀ _ <| ha hx).trans <| h₁ _ <| hb hy + +theorem image2_upperBounds_lowerBounds_subset_upperBounds_image2 : + image2 f (upperBounds s) (lowerBounds t) ⊆ upperBounds (image2 f s t) := + image2_subset_iff.2 fun _ ha _ hb ↦ + mem_upperBounds_image2_of_mem_upperBounds_of_mem_lowerBounds h₀ h₁ ha hb + +theorem image2_lowerBounds_upperBounds_subset_lowerBounds_image2 : + image2 f (lowerBounds s) (upperBounds t) ⊆ lowerBounds (image2 f s t) := + image2_subset_iff.2 fun _ ha _ hb ↦ + mem_lowerBounds_image2_of_mem_lowerBounds_of_mem_upperBounds h₀ h₁ ha hb + +theorem BddAbove.bddAbove_image2_of_bddBelow : + BddAbove s → BddBelow t → BddAbove (Set.image2 f s t) := by + rintro ⟨a, ha⟩ ⟨b, hb⟩ + exact ⟨f a b, mem_upperBounds_image2_of_mem_upperBounds_of_mem_lowerBounds h₀ h₁ ha hb⟩ + +theorem BddBelow.bddBelow_image2_of_bddAbove : + BddBelow s → BddAbove t → BddBelow (Set.image2 f s t) := by + rintro ⟨a, ha⟩ ⟨b, hb⟩ + exact ⟨f a b, mem_lowerBounds_image2_of_mem_lowerBounds_of_mem_upperBounds h₀ h₁ ha hb⟩ + +theorem IsGreatest.isGreatest_image2_of_isLeast (ha : IsGreatest s a) (hb : IsLeast t b) : + IsGreatest (Set.image2 f s t) (f a b) := + ⟨mem_image2_of_mem ha.1 hb.1, + mem_upperBounds_image2_of_mem_upperBounds_of_mem_lowerBounds h₀ h₁ ha.2 hb.2⟩ + +theorem IsLeast.isLeast_image2_of_isGreatest (ha : IsLeast s a) (hb : IsGreatest t b) : + IsLeast (Set.image2 f s t) (f a b) := + ⟨mem_image2_of_mem ha.1 hb.1, + mem_lowerBounds_image2_of_mem_lowerBounds_of_mem_upperBounds h₀ h₁ ha.2 hb.2⟩ + +end MonotoneAntitone + +section AntitoneAntitone + +variable (h₀ : ∀ b, Antitone (swap f b)) (h₁ : ∀ a, Antitone (f a)) + +include h₀ h₁ + +theorem mem_upperBounds_image2_of_mem_lowerBounds (ha : a ∈ lowerBounds s) + (hb : b ∈ lowerBounds t) : f a b ∈ upperBounds (image2 f s t) := + forall_image2_iff.2 fun _ hx _ hy => (h₀ _ <| ha hx).trans <| h₁ _ <| hb hy + +theorem mem_lowerBounds_image2_of_mem_upperBounds (ha : a ∈ upperBounds s) + (hb : b ∈ upperBounds t) : f a b ∈ lowerBounds (image2 f s t) := + forall_image2_iff.2 fun _ hx _ hy => (h₀ _ <| ha hx).trans <| h₁ _ <| hb hy + +theorem image2_upperBounds_upperBounds_subset_upperBounds_image2 : + image2 f (lowerBounds s) (lowerBounds t) ⊆ upperBounds (image2 f s t) := + image2_subset_iff.2 fun _ ha _ hb ↦ + mem_upperBounds_image2_of_mem_lowerBounds h₀ h₁ ha hb + +theorem image2_lowerBounds_lowerBounds_subset_lowerBounds_image2 : + image2 f (upperBounds s) (upperBounds t) ⊆ lowerBounds (image2 f s t) := + image2_subset_iff.2 fun _ ha _ hb ↦ + mem_lowerBounds_image2_of_mem_upperBounds h₀ h₁ ha hb + +theorem BddBelow.image2_bddAbove : BddBelow s → BddBelow t → BddAbove (Set.image2 f s t) := by + rintro ⟨a, ha⟩ ⟨b, hb⟩ + exact ⟨f a b, mem_upperBounds_image2_of_mem_lowerBounds h₀ h₁ ha hb⟩ + +theorem BddAbove.image2_bddBelow : BddAbove s → BddAbove t → BddBelow (Set.image2 f s t) := by + rintro ⟨a, ha⟩ ⟨b, hb⟩ + exact ⟨f a b, mem_lowerBounds_image2_of_mem_upperBounds h₀ h₁ ha hb⟩ + +theorem IsLeast.isGreatest_image2 (ha : IsLeast s a) (hb : IsLeast t b) : + IsGreatest (Set.image2 f s t) (f a b) := + ⟨mem_image2_of_mem ha.1 hb.1, mem_upperBounds_image2_of_mem_lowerBounds h₀ h₁ ha.2 hb.2⟩ + +theorem IsGreatest.isLeast_image2 (ha : IsGreatest s a) (hb : IsGreatest t b) : + IsLeast (Set.image2 f s t) (f a b) := + ⟨mem_image2_of_mem ha.1 hb.1, mem_lowerBounds_image2_of_mem_upperBounds h₀ h₁ ha.2 hb.2⟩ + +end AntitoneAntitone + +section AntitoneMonotone + +variable (h₀ : ∀ b, Antitone (swap f b)) (h₁ : ∀ a, Monotone (f a)) + +include h₀ h₁ + +theorem mem_upperBounds_image2_of_mem_upperBounds_of_mem_upperBounds (ha : a ∈ lowerBounds s) + (hb : b ∈ upperBounds t) : f a b ∈ upperBounds (image2 f s t) := + forall_image2_iff.2 fun _ hx _ hy => (h₀ _ <| ha hx).trans <| h₁ _ <| hb hy + +theorem mem_lowerBounds_image2_of_mem_lowerBounds_of_mem_lowerBounds (ha : a ∈ upperBounds s) + (hb : b ∈ lowerBounds t) : f a b ∈ lowerBounds (image2 f s t) := + forall_image2_iff.2 fun _ hx _ hy => (h₀ _ <| ha hx).trans <| h₁ _ <| hb hy + +theorem image2_lowerBounds_upperBounds_subset_upperBounds_image2 : + image2 f (lowerBounds s) (upperBounds t) ⊆ upperBounds (image2 f s t) := + image2_subset_iff.2 fun _ ha _ hb ↦ + mem_upperBounds_image2_of_mem_upperBounds_of_mem_upperBounds h₀ h₁ ha hb + +theorem image2_upperBounds_lowerBounds_subset_lowerBounds_image2 : + image2 f (upperBounds s) (lowerBounds t) ⊆ lowerBounds (image2 f s t) := + image2_subset_iff.2 fun _ ha _ hb ↦ + mem_lowerBounds_image2_of_mem_lowerBounds_of_mem_lowerBounds h₀ h₁ ha hb + +theorem BddBelow.bddAbove_image2_of_bddAbove : + BddBelow s → BddAbove t → BddAbove (Set.image2 f s t) := by + rintro ⟨a, ha⟩ ⟨b, hb⟩ + exact ⟨f a b, mem_upperBounds_image2_of_mem_upperBounds_of_mem_upperBounds h₀ h₁ ha hb⟩ + +theorem BddAbove.bddBelow_image2_of_bddAbove : + BddAbove s → BddBelow t → BddBelow (Set.image2 f s t) := by + rintro ⟨a, ha⟩ ⟨b, hb⟩ + exact ⟨f a b, mem_lowerBounds_image2_of_mem_lowerBounds_of_mem_lowerBounds h₀ h₁ ha hb⟩ + +theorem IsLeast.isGreatest_image2_of_isGreatest (ha : IsLeast s a) (hb : IsGreatest t b) : + IsGreatest (Set.image2 f s t) (f a b) := + ⟨mem_image2_of_mem ha.1 hb.1, + mem_upperBounds_image2_of_mem_upperBounds_of_mem_upperBounds h₀ h₁ ha.2 hb.2⟩ + +theorem IsGreatest.isLeast_image2_of_isLeast (ha : IsGreatest s a) (hb : IsLeast t b) : + IsLeast (Set.image2 f s t) (f a b) := + ⟨mem_image2_of_mem ha.1 hb.1, + mem_lowerBounds_image2_of_mem_lowerBounds_of_mem_lowerBounds h₀ h₁ ha.2 hb.2⟩ + +end AntitoneMonotone + +end Image2 + +section Prod + +variable {α β : Type*} [Preorder α] [Preorder β] + +lemma bddAbove_prod {s : Set (α × β)} : + BddAbove s ↔ BddAbove (Prod.fst '' s) ∧ BddAbove (Prod.snd '' s) := + ⟨fun ⟨p, hp⟩ ↦ ⟨⟨p.1, forall_mem_image.2 fun _q hq ↦ (hp hq).1⟩, + ⟨p.2, forall_mem_image.2 fun _q hq ↦ (hp hq).2⟩⟩, + fun ⟨⟨x, hx⟩, ⟨y, hy⟩⟩ ↦ ⟨⟨x, y⟩, fun _p hp ↦ + ⟨hx <| mem_image_of_mem _ hp, hy <| mem_image_of_mem _ hp⟩⟩⟩ + +lemma bddBelow_prod {s : Set (α × β)} : + BddBelow s ↔ BddBelow (Prod.fst '' s) ∧ BddBelow (Prod.snd '' s) := + bddAbove_prod (α := αᵒᵈ) (β := βᵒᵈ) + +lemma bddAbove_range_prod {F : ι → α × β} : + BddAbove (range F) ↔ BddAbove (range <| Prod.fst ∘ F) ∧ BddAbove (range <| Prod.snd ∘ F) := by + simp only [bddAbove_prod, ← range_comp] + +lemma bddBelow_range_prod {F : ι → α × β} : + BddBelow (range F) ↔ BddBelow (range <| Prod.fst ∘ F) ∧ BddBelow (range <| Prod.snd ∘ F) := + bddAbove_range_prod (α := αᵒᵈ) (β := βᵒᵈ) + +theorem isLUB_prod {s : Set (α × β)} (p : α × β) : + IsLUB s p ↔ IsLUB (Prod.fst '' s) p.1 ∧ IsLUB (Prod.snd '' s) p.2 := by + refine + ⟨fun H => + ⟨⟨monotone_fst.mem_upperBounds_image H.1, fun a ha => ?_⟩, + ⟨monotone_snd.mem_upperBounds_image H.1, fun a ha => ?_⟩⟩, + fun H => ⟨?_, ?_⟩⟩ + · suffices h : (a, p.2) ∈ upperBounds s from (H.2 h).1 + exact fun q hq => ⟨ha <| mem_image_of_mem _ hq, (H.1 hq).2⟩ + · suffices h : (p.1, a) ∈ upperBounds s from (H.2 h).2 + exact fun q hq => ⟨(H.1 hq).1, ha <| mem_image_of_mem _ hq⟩ + · exact fun q hq => ⟨H.1.1 <| mem_image_of_mem _ hq, H.2.1 <| mem_image_of_mem _ hq⟩ + · exact fun q hq => + ⟨H.1.2 <| monotone_fst.mem_upperBounds_image hq, + H.2.2 <| monotone_snd.mem_upperBounds_image hq⟩ + +theorem isGLB_prod {s : Set (α × β)} (p : α × β) : + IsGLB s p ↔ IsGLB (Prod.fst '' s) p.1 ∧ IsGLB (Prod.snd '' s) p.2 := + @isLUB_prod αᵒᵈ βᵒᵈ _ _ _ _ + +end Prod + + +section Pi + +variable {π : α → Type*} [∀ a, Preorder (π a)] + +lemma bddAbove_pi {s : Set (∀ a, π a)} : + BddAbove s ↔ ∀ a, BddAbove (Function.eval a '' s) := + ⟨fun ⟨f, hf⟩ a ↦ ⟨f a, forall_mem_image.2 fun _ hg ↦ hf hg a⟩, + fun h ↦ ⟨fun a ↦ (h a).some, fun _ hg a ↦ (h a).some_mem <| mem_image_of_mem _ hg⟩⟩ + +lemma bddBelow_pi {s : Set (∀ a, π a)} : + BddBelow s ↔ ∀ a, BddBelow (Function.eval a '' s) := + bddAbove_pi (π := fun a ↦ (π a)ᵒᵈ) + +lemma bddAbove_range_pi {F : ι → ∀ a, π a} : + BddAbove (range F) ↔ ∀ a, BddAbove (range fun i ↦ F i a) := by + simp only [bddAbove_pi, ← range_comp] + rfl + +lemma bddBelow_range_pi {F : ι → ∀ a, π a} : + BddBelow (range F) ↔ ∀ a, BddBelow (range fun i ↦ F i a) := + bddAbove_range_pi (π := fun a ↦ (π a)ᵒᵈ) + +theorem isLUB_pi {s : Set (∀ a, π a)} {f : ∀ a, π a} : + IsLUB s f ↔ ∀ a, IsLUB (Function.eval a '' s) (f a) := by + classical + refine + ⟨fun H a => ⟨(Function.monotone_eval a).mem_upperBounds_image H.1, fun b hb => ?_⟩, fun H => + ⟨?_, ?_⟩⟩ + · suffices h : Function.update f a b ∈ upperBounds s from Function.update_same a b f ▸ H.2 h a + exact fun g hg => le_update_iff.2 ⟨hb <| mem_image_of_mem _ hg, fun i _ => H.1 hg i⟩ + · exact fun g hg a => (H a).1 (mem_image_of_mem _ hg) + · exact fun g hg a => (H a).2 ((Function.monotone_eval a).mem_upperBounds_image hg) + +theorem isGLB_pi {s : Set (∀ a, π a)} {f : ∀ a, π a} : + IsGLB s f ↔ ∀ a, IsGLB (Function.eval a '' s) (f a) := + @isLUB_pi α (fun a => (π a)ᵒᵈ) _ s f + +end Pi + +theorem IsGLB.of_image [Preorder α] [Preorder β] {f : α → β} (hf : ∀ {x y}, f x ≤ f y ↔ x ≤ y) + {s : Set α} {x : α} (hx : IsGLB (f '' s) (f x)) : IsGLB s x := + ⟨fun _ hy => hf.1 <| hx.1 <| mem_image_of_mem _ hy, fun _ hy => + hf.1 <| hx.2 <| Monotone.mem_lowerBounds_image (fun _ _ => hf.2) hy⟩ + +theorem IsLUB.of_image [Preorder α] [Preorder β] {f : α → β} (hf : ∀ {x y}, f x ≤ f y ↔ x ≤ y) + {s : Set α} {x : α} (hx : IsLUB (f '' s) (f x)) : IsLUB s x := + ⟨fun _ hy => hf.1 <| hx.1 <| mem_image_of_mem _ hy, fun _ hy => + hf.1 <| hx.2 <| Monotone.mem_upperBounds_image (fun _ _ => hf.2) hy⟩ + +lemma BddAbove.range_mono [Preorder β] {f : α → β} (g : α → β) (h : ∀ a, f a ≤ g a) + (hbdd : BddAbove (range g)) : BddAbove (range f) := by + obtain ⟨C, hC⟩ := hbdd + use C + rintro - ⟨x, rfl⟩ + exact (h x).trans (hC <| mem_range_self x) + +lemma BddBelow.range_mono [Preorder β] (f : α → β) {g : α → β} (h : ∀ a, f a ≤ g a) + (hbdd : BddBelow (range f)) : BddBelow (range g) := + BddAbove.range_mono (β := βᵒᵈ) f h hbdd + +lemma BddAbove.range_comp {γ : Type*} [Preorder β] [Preorder γ] {f : α → β} {g : β → γ} + (hf : BddAbove (range f)) (hg : Monotone g) : BddAbove (range (fun x => g (f x))) := by + change BddAbove (range (g ∘ f)) + simpa only [Set.range_comp] using hg.map_bddAbove hf + +lemma BddBelow.range_comp {γ : Type*} [Preorder β] [Preorder γ] {f : α → β} {g : β → γ} + (hf : BddBelow (range f)) (hg : Monotone g) : BddBelow (range (fun x => g (f x))) := by + change BddBelow (range (g ∘ f)) + simpa only [Set.range_comp] using hg.map_bddBelow hf diff --git a/Mathlib/Order/Bounds/OrderIso.lean b/Mathlib/Order/Bounds/OrderIso.lean index 913266a83c0cc..d3fda94252a64 100644 --- a/Mathlib/Order/Bounds/OrderIso.lean +++ b/Mathlib/Order/Bounds/OrderIso.lean @@ -3,7 +3,7 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Yury Kudryashov -/ -import Mathlib.Order.Bounds.Basic +import Mathlib.Order.Bounds.Image import Mathlib.Order.Hom.Set /-! diff --git a/Mathlib/Order/CompleteLattice.lean b/Mathlib/Order/CompleteLattice.lean index 2d47e7b97ad07..b2c06f1f0a3d0 100644 --- a/Mathlib/Order/CompleteLattice.lean +++ b/Mathlib/Order/CompleteLattice.lean @@ -5,6 +5,7 @@ Authors: Johannes Hölzl -/ import Mathlib.Data.Bool.Set import Mathlib.Data.Nat.Set +import Mathlib.Data.Set.NAry import Mathlib.Data.Set.Prod import Mathlib.Data.ULift import Mathlib.Order.Bounds.Basic diff --git a/Mathlib/Order/GaloisConnection.lean b/Mathlib/Order/GaloisConnection.lean index 7420bcaf290d1..5995f271ba962 100644 --- a/Mathlib/Order/GaloisConnection.lean +++ b/Mathlib/Order/GaloisConnection.lean @@ -6,7 +6,7 @@ Authors: Johannes Hölzl import Mathlib.Order.CompleteLattice import Mathlib.Order.Synonym import Mathlib.Order.Hom.Set -import Mathlib.Order.Bounds.Basic +import Mathlib.Order.Bounds.Image /-! # Galois connections, insertions and coinsertions From e06fe3e853532f908b0e52e0e65db18e4dd65a76 Mon Sep 17 00:00:00 2001 From: FR Date: Sat, 19 Oct 2024 16:03:52 +0000 Subject: [PATCH 186/505] =?UTF-8?q?chore:=20do=20not=20use=20`=C2=B7=20?= =?UTF-8?q?=E2=89=88=20=C2=B7`=20in=20`Quotient.mk=5Fout`=20(#17940)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- Mathlib/Data/Quot.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Data/Quot.lean b/Mathlib/Data/Quot.lean index 8dfbafe4dbf99..b9ff59544232d 100644 --- a/Mathlib/Data/Quot.lean +++ b/Mathlib/Data/Quot.lean @@ -355,7 +355,7 @@ noncomputable def Quotient.out {s : Setoid α} : Quotient s → α := theorem Quotient.out_eq {s : Setoid α} (q : Quotient s) : ⟦q.out⟧ = q := Quot.out_eq q -theorem Quotient.mk_out {s : Setoid α} (a : α) : (⟦a⟧ : Quotient s).out ≈ a := +theorem Quotient.mk_out {s : Setoid α} (a : α) : s (⟦a⟧ : Quotient s).out a := Quotient.exact (Quotient.out_eq _) theorem Quotient.mk_eq_iff_out {s : Setoid α} {x : α} {y : Quotient s} : From 2c5ee0d7efdb723afdf718540ddfabaa477015fc Mon Sep 17 00:00:00 2001 From: Yongle Hu Date: Sat, 19 Oct 2024 18:37:47 +0000 Subject: [PATCH 187/505] feat(RingTheory/Ideal/Over): define a class for an ideal lying over an ideal (#17415) Define `class` of an ideal lies over an ideal. Co-authored-by: Yongle Hu @mbkybky Co-authored-by: Jiedong Jiang @jjdishere Co-authored-by: Hu Yongle <2065545849@qq.com> Co-authored-by: Yongle Hu <2065545849@qq.com> --- Mathlib/RingTheory/FiniteType.lean | 5 + Mathlib/RingTheory/Ideal/Maps.lean | 16 +++ Mathlib/RingTheory/Ideal/Over.lean | 126 +++++++++++++++++- .../IsIntegralClosure/Basic.lean | 25 ++++ Mathlib/RingTheory/Noetherian.lean | 8 +- 5 files changed, 176 insertions(+), 4 deletions(-) diff --git a/Mathlib/RingTheory/FiniteType.lean b/Mathlib/RingTheory/FiniteType.lean index 8eb4df96ddf80..97c79c644e56a 100644 --- a/Mathlib/RingTheory/FiniteType.lean +++ b/Mathlib/RingTheory/FiniteType.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin -/ import Mathlib.RingTheory.Adjoin.Tower +import Mathlib.RingTheory.Ideal.QuotientOperations /-! # Finiteness conditions in commutative algebra @@ -111,6 +112,10 @@ theorem trans [Algebra S A] [IsScalarTower R S A] (hRS : FiniteType R S) (hSA : FiniteType R A := ⟨fg_trans' hRS.1 hSA.1⟩ +instance quotient (R : Type*) {S : Type*} [CommSemiring R] [CommRing S] [Algebra R S] (I : Ideal S) + [h : Algebra.FiniteType R S] : Algebra.FiniteType R (S ⧸ I) := + Algebra.FiniteType.trans h inferInstance + /-- An algebra is finitely generated if and only if it is a quotient of a free algebra whose variables are indexed by a finset. -/ theorem iff_quotient_freeAlgebra : diff --git a/Mathlib/RingTheory/Ideal/Maps.lean b/Mathlib/RingTheory/Ideal/Maps.lean index faa5fc206c934..80ddaa2813840 100644 --- a/Mathlib/RingTheory/Ideal/Maps.lean +++ b/Mathlib/RingTheory/Ideal/Maps.lean @@ -846,3 +846,19 @@ def idealMap (I : Ideal R) : I →ₗ[R] I.map (algebraMap R S) := (fun _ ↦ Ideal.mem_map_of_mem _) end Algebra + +namespace NoZeroSMulDivisors + +theorem of_ker_algebraMap_eq_bot (R A : Type*) [CommRing R] [Semiring A] [Algebra R A] + [NoZeroDivisors A] (h : RingHom.ker (algebraMap R A) = ⊥) : NoZeroSMulDivisors R A := + of_algebraMap_injective ((RingHom.injective_iff_ker_eq_bot _).mpr h) + +theorem ker_algebraMap_eq_bot (R A : Type*) [CommRing R] [Ring A] [Nontrivial A] [Algebra R A] + [NoZeroSMulDivisors R A] : RingHom.ker (algebraMap R A) = ⊥ := + (RingHom.injective_iff_ker_eq_bot _).mp (algebraMap_injective R A) + +theorem iff_ker_algebraMap_eq_bot {R A : Type*} [CommRing R] [Ring A] [IsDomain A] [Algebra R A] : + NoZeroSMulDivisors R A ↔ RingHom.ker (algebraMap R A) = ⊥ := + iff_algebraMap_injective.trans (RingHom.injective_iff_ker_eq_bot (algebraMap R A)) + +end NoZeroSMulDivisors diff --git a/Mathlib/RingTheory/Ideal/Over.lean b/Mathlib/RingTheory/Ideal/Over.lean index f8586e0f2559a..bbc4a002eeb87 100644 --- a/Mathlib/RingTheory/Ideal/Over.lean +++ b/Mathlib/RingTheory/Ideal/Over.lean @@ -3,7 +3,6 @@ Copyright (c) 2020 Anne Baanen. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Anne Baanen -/ -import Mathlib.RingTheory.Algebraic import Mathlib.RingTheory.Localization.AtPrime import Mathlib.RingTheory.Localization.Integral @@ -429,4 +428,129 @@ lemma map_eq_top_iff {R S} [CommRing R] [CommRing S] end IsDomain +section ideal_liesOver + +section Semiring + +variable (A : Type*) [CommSemiring A] {B : Type*} [Semiring B] [Algebra A B] + (P : Ideal B) (p : Ideal A) + +/-- The ideal obtained by pulling back the ideal `P` from `B` to `A`. -/ +abbrev under : Ideal A := Ideal.comap (algebraMap A B) P + +theorem under_def : P.under A = Ideal.comap (algebraMap A B) P := rfl + +instance IsPrime.under [hP : P.IsPrime] : (P.under A).IsPrime := + hP.comap (algebraMap A B) + +variable {A} + +/-- `P` lies over `p` if `p` is the preimage of `P` of the `algebraMap`. -/ +class LiesOver : Prop where + over : p = P.under A + +instance over_under : P.LiesOver (P.under A) where over := rfl + +theorem over_def [P.LiesOver p] : p = P.under A := LiesOver.over + +theorem mem_of_liesOver [P.LiesOver p] (x : A) : x ∈ p ↔ algebraMap A B x ∈ P := by + rw [P.over_def p] + rfl + +end Semiring + +section CommSemiring + +variable {A : Type*} [CommSemiring A] {B : Type*} [CommSemiring B] {C : Type*} [Semiring C] + [Algebra A B] [Algebra B C] [Algebra A C] [IsScalarTower A B C] + (𝔓 : Ideal C) (P : Ideal B) (p : Ideal A) + +@[simp] +theorem under_under : (𝔓.under B).under A = 𝔓.under A := by + simp_rw [comap_comap, ← IsScalarTower.algebraMap_eq] + +theorem LiesOver.trans [𝔓.LiesOver P] [P.LiesOver p] : 𝔓.LiesOver p where + over := by rw [P.over_def p, 𝔓.over_def P, under_under] + +theorem LiesOver.tower_bot [hp : 𝔓.LiesOver p] [hP : 𝔓.LiesOver P] : P.LiesOver p where + over := by rw [𝔓.over_def p, 𝔓.over_def P, under_under] + +variable (B) + +instance under_liesOver_of_liesOver [𝔓.LiesOver p] : (𝔓.under B).LiesOver p := + LiesOver.tower_bot 𝔓 (𝔓.under B) p + +end CommSemiring + +section CommRing + +namespace Quotient + +variable (R : Type*) [CommSemiring R] {A B : Type*} [CommRing A] [CommRing B] [Algebra A B] + [Algebra R A] [Algebra R B] [IsScalarTower R A B] (P : Ideal B) (p : Ideal A) [P.LiesOver p] + +/-- If `P` lies over `p`, then canonically `B ⧸ P` is a `A ⧸ p`-algebra. -/ +instance algebraOfLiesOver : Algebra (A ⧸ p) (B ⧸ P) := + algebraQuotientOfLEComap (le_of_eq (P.over_def p)) + +instance isScalarTower_of_liesOver : IsScalarTower R (A ⧸ p) (B ⧸ P) := + IsScalarTower.of_algebraMap_eq' <| + congrArg (algebraMap B (B ⧸ P)).comp (IsScalarTower.algebraMap_eq R A B) + +/-- `B ⧸ P` is a finite `A ⧸ p`-module if `B` is a finite `A`-module. -/ +instance module_finite_of_liesOver [Module.Finite A B] : Module.Finite (A ⧸ p) (B ⧸ P) := + Module.Finite.of_restrictScalars_finite A (A ⧸ p) (B ⧸ P) + +example [Module.Finite A B] : Module.Finite (A ⧸ P.under A) (B ⧸ P) := inferInstance + +/-- `B ⧸ P` is a finitely generated `A ⧸ p`-algebra if `B` is a finitely generated `A`-algebra. -/ +instance algebra_finiteType_of_liesOver [Algebra.FiniteType A B] : + Algebra.FiniteType (A ⧸ p) (B ⧸ P) := + Algebra.FiniteType.of_restrictScalars_finiteType A (A ⧸ p) (B ⧸ P) + +/-- `B ⧸ P` is a Noetherian `A ⧸ p`-module if `B` is a Noetherian `A`-module. -/ +instance isNoetherian_of_liesOver [IsNoetherian A B] : IsNoetherian (A ⧸ p) (B ⧸ P) := + isNoetherian_of_tower A inferInstance + +theorem algebraMap_injective_of_liesOver : + Function.Injective (algebraMap (A ⧸ p) (B ⧸ P)) := by + rintro ⟨a⟩ ⟨b⟩ hab + apply Quotient.eq.mpr ((mem_of_liesOver P p (a - b)).mpr _) + rw [RingHom.map_sub] + exact Quotient.eq.mp hab + +instance [P.IsPrime] : NoZeroSMulDivisors (A ⧸ p) (B ⧸ P) := + NoZeroSMulDivisors.of_algebraMap_injective (Quotient.algebraMap_injective_of_liesOver P p) + +end Quotient + +end CommRing + +section IsIntegral + +variable {A : Type*} [CommRing A] {B : Type*} [CommRing B] [Algebra A B] [Algebra.IsIntegral A B] + (P : Ideal B) (p : Ideal A) [P.LiesOver p] + +variable (A) in +/-- If `B` is an integral `A`-algebra, `P` is a maximal ideal of `B`, then the pull back of + `P` is also a maximal ideal of `A`. -/ +instance IsMaximal.under [P.IsMaximal] : (P.under A).IsMaximal := + isMaximal_comap_of_isIntegral_of_isMaximal P + +theorem IsMaximal.of_liesOver_isMaximal [hpm : p.IsMaximal] [P.IsPrime] : P.IsMaximal := by + rw [P.over_def p] at hpm + exact isMaximal_of_isIntegral_of_isMaximal_comap P hpm + +theorem IsMaximal.of_isMaximal_liesOver [P.IsMaximal] : p.IsMaximal := by + rw [P.over_def p] + exact isMaximal_comap_of_isIntegral_of_isMaximal P + +/-- `B ⧸ P` is an integral `A ⧸ p`-algebra if `B` is a integral `A`-algebra. -/ +instance Quotient.algebra_isIntegral_of_liesOver : Algebra.IsIntegral (A ⧸ p) (B ⧸ P) := + Algebra.IsIntegral.tower_top A + +end IsIntegral + +end ideal_liesOver + end Ideal diff --git a/Mathlib/RingTheory/IntegralClosure/IsIntegralClosure/Basic.lean b/Mathlib/RingTheory/IntegralClosure/IsIntegralClosure/Basic.lean index 59406a95b8ec6..eabc18718c598 100644 --- a/Mathlib/RingTheory/IntegralClosure/IsIntegralClosure/Basic.lean +++ b/Mathlib/RingTheory/IntegralClosure/IsIntegralClosure/Basic.lean @@ -558,6 +558,18 @@ nonrec theorem RingHom.IsIntegral.tower_bot (hg : Function.Injective g) haveI : IsScalarTower R S T := IsScalarTower.of_algebraMap_eq fun _ ↦ rfl fun x ↦ IsIntegral.tower_bot hg (hfg (g x)) +variable (T) in +/-- Let `T / S / R` be a tower of algebras, `T` is non-trivial and is a torsion free `S`-module, + then if `T` is an integral `R`-algebra, then `S` is an integral `R`-algebra. -/ +theorem Algebra.IsIntegral.tower_bot [Algebra R S] [Algebra R T] [Algebra S T] + [NoZeroSMulDivisors S T] [Nontrivial T] [IsScalarTower R S T] + [h : Algebra.IsIntegral R T] : Algebra.IsIntegral R S where + isIntegral := by + apply RingHom.IsIntegral.tower_bot (algebraMap R S) (algebraMap S T) + (NoZeroSMulDivisors.algebraMap_injective S T) + rw [← IsScalarTower.algebraMap_eq R S T] + exact h.isIntegral + theorem IsIntegral.tower_bot_of_field {R A B : Type*} [CommRing R] [Field A] [CommRing B] [Nontrivial B] [Algebra R A] [Algebra A B] [Algebra R B] [IsScalarTower R A B] {x : A} (h : IsIntegral R (algebraMap A B x)) : IsIntegral R x := @@ -571,6 +583,16 @@ theorem RingHom.isIntegralElem.of_comp {x : T} (h : (g.comp f).IsIntegralElem x) theorem RingHom.IsIntegral.tower_top (h : (g.comp f).IsIntegral) : g.IsIntegral := fun x ↦ RingHom.isIntegralElem.of_comp f g (h x) +variable (R) in +/-- Let `T / S / R` be a tower of algebras, `T` is an integral `R`-algebra, then it is integral + as an `S`-algebra. -/ +theorem Algebra.IsIntegral.tower_top [Algebra R S] [Algebra R T] [Algebra S T] [IsScalarTower R S T] + [h : Algebra.IsIntegral R T] : Algebra.IsIntegral S T where + isIntegral := by + apply RingHom.IsIntegral.tower_top (algebraMap R S) (algebraMap S T) + rw [← IsScalarTower.algebraMap_eq R S T] + exact h.isIntegral + theorem RingHom.IsIntegral.quotient {I : Ideal S} (hf : f.IsIntegral) : (Ideal.quotientMap I f le_rfl).IsIntegral := by rintro ⟨x⟩ @@ -578,6 +600,9 @@ theorem RingHom.IsIntegral.quotient {I : Ideal S} (hf : f.IsIntegral) : refine ⟨p.map (Ideal.Quotient.mk _), p_monic.map _, ?_⟩ simpa only [hom_eval₂, eval₂_map] using congr_arg (Ideal.Quotient.mk I) hpx +instance {I : Ideal A} [Algebra.IsIntegral R A] : Algebra.IsIntegral R (A ⧸ I) := + Algebra.IsIntegral.trans A + instance Algebra.IsIntegral.quotient {I : Ideal A} [Algebra.IsIntegral R A] : Algebra.IsIntegral (R ⧸ I.comap (algebraMap R A)) (A ⧸ I) := ⟨RingHom.IsIntegral.quotient (algebraMap R A) Algebra.IsIntegral.isIntegral⟩ diff --git a/Mathlib/RingTheory/Noetherian.lean b/Mathlib/RingTheory/Noetherian.lean index 903c8862e2b78..411b9fdf6c5a4 100644 --- a/Mathlib/RingTheory/Noetherian.lean +++ b/Mathlib/RingTheory/Noetherian.lean @@ -112,9 +112,11 @@ theorem isNoetherian_of_surjective (f : M →ₗ[R] P) (hf : LinearMap.range f = variable {M} -instance isNoetherian_quotient {R} [Ring R] {M} [AddCommGroup M] [Module R M] - (N : Submodule R M) [IsNoetherian R M] : IsNoetherian R (M ⧸ N) := - isNoetherian_of_surjective _ _ (LinearMap.range_eq_top.mpr N.mkQ_surjective) +instance isNoetherian_quotient {A M : Type*} [Ring A] [AddCommGroup M] [SMul R A] [Module R M] + [Module A M] [IsScalarTower R A M] (N : Submodule A M) [IsNoetherian R M] : + IsNoetherian R (M ⧸ N) := + isNoetherian_of_surjective M ((Submodule.mkQ N).restrictScalars R) <| + LinearMap.range_eq_top.mpr N.mkQ_surjective @[deprecated (since := "2024-04-27"), nolint defLemma] alias Submodule.Quotient.isNoetherian := isNoetherian_quotient From 37de4d77648a7d69b80dfc65ad9ef35e65696d2f Mon Sep 17 00:00:00 2001 From: Jakob von Raumer Date: Sat, 19 Oct 2024 19:08:11 +0000 Subject: [PATCH 188/505] feat(CategoryTheory/KanExtensions): Evaluating `lan` twice amounts to taking a colimit (#17531) --- .../Functor/KanExtension/Adjunction.lean | 44 +++++++++++++++++ .../Functor/KanExtension/Pointwise.lean | 48 +++++++++++++++++++ 2 files changed, 92 insertions(+) diff --git a/Mathlib/CategoryTheory/Functor/KanExtension/Adjunction.lean b/Mathlib/CategoryTheory/Functor/KanExtension/Adjunction.lean index fe8433c905fb7..274273b8794c3 100644 --- a/Mathlib/CategoryTheory/Functor/KanExtension/Adjunction.lean +++ b/Mathlib/CategoryTheory/Functor/KanExtension/Adjunction.lean @@ -56,6 +56,28 @@ noncomputable def isPointwiseLeftKanExtensionLanUnit (LeftExtension.mk _ (L.lanUnit.app F)).IsPointwiseLeftKanExtension := isPointwiseLeftKanExtensionOfIsLeftKanExtension (F := F) _ (L.lanUnit.app F) +/-- If a left Kan extension is pointwise, then evaluating it at an object is isomorphic to +taking a colimit. -/ +noncomputable def lanObjObjIsoColimit (F : C ⥤ H) [HasPointwiseLeftKanExtension L F] (X : D) : + (L.lan.obj F).obj X ≅ Limits.colimit (CostructuredArrow.proj L X ⋙ F) := + LeftExtension.IsPointwiseLeftKanExtensionAt.isoColimit (F := F) + (isPointwiseLeftKanExtensionLanUnit L F X) + +@[reassoc (attr := simp)] +lemma ι_lanObjObjIsoColimit_inv + (F : C ⥤ H) [HasPointwiseLeftKanExtension L F] (X : D) (f : CostructuredArrow L X) : + Limits.colimit.ι _ f ≫ (L.lanObjObjIsoColimit F X).inv = + (L.lanUnit.app F).app f.left ≫ (L.lan.obj F).map f.hom := by + simp [lanObjObjIsoColimit, lanUnit] + +@[reassoc (attr := simp)] +lemma ι_lanObjObjIsoColimit_hom + (F : C ⥤ H) [HasPointwiseLeftKanExtension L F] (X : D) (f : CostructuredArrow L X) : + (L.lanUnit.app F).app f.left ≫ (L.lan.obj F).map f.hom ≫ (L.lanObjObjIsoColimit F X).hom = + Limits.colimit.ι (CostructuredArrow.proj L X ⋙ F) f := + LeftExtension.IsPointwiseLeftKanExtensionAt.ι_isoColimit_hom (F := F) + (isPointwiseLeftKanExtensionLanUnit L F X) f + variable (H) in /-- The left Kan extension functor `L.Lan` is left adjoint to the precomposition by `L`. -/ @@ -162,6 +184,28 @@ noncomputable def isPointwiseRightKanExtensionRanCounit (RightExtension.mk _ (L.ranCounit.app F)).IsPointwiseRightKanExtension := isPointwiseRightKanExtensionOfIsRightKanExtension (F := F) _ (L.ranCounit.app F) +/-- If a right Kan extension is pointwise, then evaluating it at an object is isomorphic to +taking a limit. -/ +noncomputable def ranObjObjIsoLimit (F : C ⥤ H) [HasPointwiseRightKanExtension L F] (X : D) : + (L.ran.obj F).obj X ≅ Limits.limit (StructuredArrow.proj X L ⋙ F) := + RightExtension.IsPointwiseRightKanExtensionAt.isoLimit (F := F) + (isPointwiseRightKanExtensionRanCounit L F X) + +@[reassoc (attr := simp)] +lemma ranObjObjIsoLimit_hom_π + (F : C ⥤ H) [HasPointwiseRightKanExtension L F] (X : D) (f : StructuredArrow X L) : + (L.ranObjObjIsoLimit F X).hom ≫ Limits.limit.π _ f = + (L.ran.obj F).map f.hom ≫ (L.ranCounit.app F).app f.right := by + simp [ranObjObjIsoLimit, ran, ranCounit] + +@[reassoc (attr := simp)] +lemma ranObjObjIsoLimit_inv_π + (F : C ⥤ H) [HasPointwiseRightKanExtension L F] (X : D) (f : StructuredArrow X L) : + (L.ranObjObjIsoLimit F X).inv ≫ (L.ran.obj F).map f.hom ≫ (L.ranCounit.app F).app f.right = + Limits.limit.π _ f := + RightExtension.IsPointwiseRightKanExtensionAt.isoLimit_inv_π (F := F) + (isPointwiseRightKanExtensionRanCounit L F X) f + variable (H) in /-- The right Kan extension functor `L.ran` is right adjoint to the precomposition by `L`. -/ diff --git a/Mathlib/CategoryTheory/Functor/KanExtension/Pointwise.lean b/Mathlib/CategoryTheory/Functor/KanExtension/Pointwise.lean index 9d4cb7a6c2ea1..a1a044d1a3846 100644 --- a/Mathlib/CategoryTheory/Functor/KanExtension/Pointwise.lean +++ b/Mathlib/CategoryTheory/Functor/KanExtension/Pointwise.lean @@ -98,6 +98,30 @@ lemma IsPointwiseLeftKanExtensionAt.isIso_hom_app IsIso (E.hom.app X) := by simpa using h.isIso_ι_app_of_isTerminal _ CostructuredArrow.mkIdTerminal +namespace IsPointwiseLeftKanExtensionAt + +variable {E} {Y : D} (h : E.IsPointwiseLeftKanExtensionAt Y) + [HasColimit (CostructuredArrow.proj L Y ⋙ F)] + +/-- A pointwise left Kan extension of `F` along `L` applied to an object `Y` is isomorphic to +`colimit (CostructuredArrow.proj L Y ⋙ F)`. -/ +noncomputable def isoColimit : + E.right.obj Y ≅ colimit (CostructuredArrow.proj L Y ⋙ F) := + h.coconePointUniqueUpToIso (colimit.isColimit _) + +@[reassoc (attr := simp)] +lemma ι_isoColimit_inv (g : CostructuredArrow L Y) : + colimit.ι _ g ≫ h.isoColimit.inv = E.hom.app g.left ≫ E.right.map g.hom := + IsColimit.comp_coconePointUniqueUpToIso_inv _ _ _ + +@[reassoc (attr := simp)] +lemma ι_isoColimit_hom (g : CostructuredArrow L Y) : + E.hom.app g.left ≫ E.right.map g.hom ≫ h.isoColimit.hom = + colimit.ι (CostructuredArrow.proj L Y ⋙ F) g := by + simpa using h.comp_coconePointUniqueUpToIso_hom (colimit.isColimit _) g + +end IsPointwiseLeftKanExtensionAt + /-- A left extension `E : LeftExtension L F` is a pointwise left Kan extension when it is a pointwise left Kan extension at any object. -/ abbrev IsPointwiseLeftKanExtension := ∀ (Y : D), E.IsPointwiseLeftKanExtensionAt Y @@ -211,6 +235,30 @@ lemma IsPointwiseRightKanExtensionAt.isIso_hom_app IsIso (E.hom.app X) := by simpa using h.isIso_π_app_of_isInitial _ StructuredArrow.mkIdInitial +namespace IsPointwiseRightKanExtensionAt + +variable {E} {Y : D} (h : E.IsPointwiseRightKanExtensionAt Y) + [HasLimit (StructuredArrow.proj Y L ⋙ F)] + +/-- A pointwise right Kan extension of `F` along `L` applied to an object `Y` is isomorphic to +`limit (StructuredArrow.proj Y L ⋙ F)`. -/ +noncomputable def isoLimit : + E.left.obj Y ≅ limit (StructuredArrow.proj Y L ⋙ F) := + h.conePointUniqueUpToIso (limit.isLimit _) + +@[reassoc (attr := simp)] +lemma isoLimit_hom_π (g : StructuredArrow Y L) : + h.isoLimit.hom ≫ limit.π _ g = E.left.map g.hom ≫ E.hom.app g.right := + IsLimit.conePointUniqueUpToIso_hom_comp _ _ _ + +@[reassoc (attr := simp)] +lemma isoLimit_inv_π (g : StructuredArrow Y L) : + h.isoLimit.inv ≫ E.left.map g.hom ≫ E.hom.app g.right = + limit.π (StructuredArrow.proj Y L ⋙ F) g := by + simpa using h.conePointUniqueUpToIso_inv_comp (limit.isLimit _) g + +end IsPointwiseRightKanExtensionAt + /-- A right extension `E : RightExtension L F` is a pointwise right Kan extension when it is a pointwise right Kan extension at any object. -/ abbrev IsPointwiseRightKanExtension := ∀ (Y : D), E.IsPointwiseRightKanExtensionAt Y From 19e2abfdd236a1feb36b33a6e219be17097d587f Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Sat, 19 Oct 2024 19:20:12 +0000 Subject: [PATCH 189/505] feat (AlgebraicGeometry/Modules): stalks of the sheaf `M^~` (#14247) Construct the isomorphism of $R$-modules from the stalk of the sheaf $M^{\~}$ on $Spec(R)$ to the localization $M_x$ as an $R$-module, where $R$ is a commutative ring and $M$ is an $R$-module. ## Main definition * `ModuleCat.tilde.stalkIso`: The isomorphism of `R`-modules from the stalk of `M^~` at `x` to the localization of `M` at the prime ideal corresponding to `x`. ## Technical note To get the `R`-module structure on the stalks on `M^~`, we had to define `ModuleCat.tildeInModuleCat`, which is `M^~` seen as sheaf of `R`-modules. We get it by applying a forgetful functor to `ModuleCat.tilde M`. Co-authored-by: Weihong Xu Co-authored-by: Sophie Morel Co-authored-by: Amelia Livingston This contribution was created as part of the AIM workshop "Formalizing algebraic geometry" in June 2024. Co-authored-by: morel Co-authored-by: Weihong Xu --- Mathlib/AlgebraicGeometry/Modules/Tilde.lean | 140 +++++++++++++++++++ 1 file changed, 140 insertions(+) diff --git a/Mathlib/AlgebraicGeometry/Modules/Tilde.lean b/Mathlib/AlgebraicGeometry/Modules/Tilde.lean index 7fd34c2070d7a..d14b580db3d1a 100644 --- a/Mathlib/AlgebraicGeometry/Modules/Tilde.lean +++ b/Mathlib/AlgebraicGeometry/Modules/Tilde.lean @@ -22,6 +22,8 @@ such that `M^~(U)` is the set of dependent functions that are locally fractions. * `ModuleCat.tildeInType` : `M^~` as a sheaf of types groups. * `ModuleCat.tilde` : `M^~` as a sheaf of `𝒪_{Spec R}`-modules. +* `ModuleCat.tilde.stalkIso`: The isomorphism of `R`-modules from the stalk of `M^~` at `x` to +the localization of `M` at the prime ideal corresponding to `x`. ## Technical note @@ -252,6 +254,144 @@ noncomputable def localizationToStalk (x : PrimeSpectrum.Top R) : (TopCat.Presheaf.stalk (tildeInModuleCat M) x) := LocalizedModule.lift _ (toStalk M x) <| isUnit_toStalk M x + +/-- The ring homomorphism that takes a section of the structure sheaf of `R` on the open set `U`, +implemented as a subtype of dependent functions to localizations at prime ideals, and evaluates +the section on the point corresponding to a given prime ideal. -/ +def openToLocalization (U : Opens (PrimeSpectrum R)) (x : PrimeSpectrum R) (hx : x ∈ U) : + (tildeInModuleCat M).obj (op U) ⟶ + ModuleCat.of R (LocalizedModule x.asIdeal.primeCompl M) where + toFun s := (s.1 ⟨x, hx⟩ : _) + map_add' _ _ := rfl + map_smul' _ _ := rfl + +/-- +The morphism of `R`-modules from the stalk of `M^~` at `x` to the localization of `M` at the +prime ideal of `R` corresponding to `x`. +-/ +noncomputable def stalkToFiberLinearMap (x : PrimeSpectrum.Top R) : + (tildeInModuleCat M).stalk x ⟶ + ModuleCat.of R (LocalizedModule x.asIdeal.primeCompl M) := + Limits.colimit.desc ((OpenNhds.inclusion x).op ⋙ (tildeInModuleCat M)) + { pt := _ + ι := + { app := fun U => openToLocalization M ((OpenNhds.inclusion _).obj U.unop) x U.unop.2 } } + +@[simp] +theorem germ_comp_stalkToFiberLinearMap (U : Opens (PrimeSpectrum.Top R)) (x) (hx : x ∈ U) : + (tildeInModuleCat M).germ U x hx ≫ stalkToFiberLinearMap M x = + openToLocalization M U x hx := + Limits.colimit.ι_desc _ _ + +@[simp] +theorem stalkToFiberLinearMap_germ (U : Opens (PrimeSpectrum.Top R)) (x : PrimeSpectrum.Top R) + (hx : x ∈ U) (s : (tildeInModuleCat M).1.obj (op U)) : + stalkToFiberLinearMap M x + (TopCat.Presheaf.germ (tildeInModuleCat M) U x hx s) = (s.1 ⟨x, hx⟩ : _) := + DFunLike.ext_iff.1 (germ_comp_stalkToFiberLinearMap M U x hx) s + +@[reassoc (attr := simp), elementwise (attr := simp)] +theorem toOpen_germ (U : Opens (PrimeSpectrum.Top R)) (x) (hx : x ∈ U) : + toOpen M U ≫ M.tildeInModuleCat.germ U x hx = toStalk M x := by + rw [← toOpen_res M ⊤ U (homOfLE le_top : U ⟶ ⊤), Category.assoc, Presheaf.germ_res]; rfl + +@[reassoc (attr := simp)] +theorem toStalk_comp_stalkToFiberLinearMap (x : PrimeSpectrum.Top R) : + toStalk M x ≫ stalkToFiberLinearMap M x = + LocalizedModule.mkLinearMap x.asIdeal.primeCompl M := by + rw [toStalk, Category.assoc, germ_comp_stalkToFiberLinearMap]; rfl + +theorem stalkToFiberLinearMap_toStalk (x : PrimeSpectrum.Top R) (m : M) : + (stalkToFiberLinearMap M x) (toStalk M x m) = + LocalizedModule.mk m 1 := + LinearMap.ext_iff.1 (toStalk_comp_stalkToFiberLinearMap M x) _ + +/-- +If `U` is an open subset of `Spec R`, `m` is an element of `M` and `r` is an element of `R` +that is invertible on `U` (i.e. does not belong to any prime ideal corresponding to a point +in `U`), this is `m / r` seen as a section of `M^~` over `U`. +-/ +def const (m : M) (r : R) (U : Opens (PrimeSpectrum.Top R)) + (hu : ∀ x ∈ U, r ∈ (x : PrimeSpectrum.Top R).asIdeal.primeCompl) : + (tildeInModuleCat M).obj (op U) := + ⟨fun x => LocalizedModule.mk m ⟨r, hu x x.2⟩, fun x => + ⟨U, x.2, 𝟙 _, m, r, fun y => ⟨hu _ y.2, by + simpa only [LocalizedModule.mkLinearMap_apply, LocalizedModule.smul'_mk, + LocalizedModule.mk_eq] using ⟨1, by simp⟩⟩⟩⟩ + +@[simp] +theorem const_apply (m : M) (r : R) (U : Opens (PrimeSpectrum.Top R)) + (hu : ∀ x ∈ U, r ∈ (x : PrimeSpectrum.Top R).asIdeal.primeCompl) (x : U) : + (const M m r U hu).1 x = LocalizedModule.mk m ⟨r, hu x x.2⟩ := + rfl + +theorem exists_const (U) (s : (tildeInModuleCat M).obj (op U)) (x : PrimeSpectrum.Top R) + (hx : x ∈ U) : + ∃ (V : Opens (PrimeSpectrum.Top R)) (_ : x ∈ V) (i : V ⟶ U) (f : M) (g : R) (hg : _), + const M f g V hg = (tildeInModuleCat M).map i.op s := + let ⟨V, hxV, iVU, f, g, hfg⟩ := s.2 ⟨x, hx⟩ + ⟨V, hxV, iVU, f, g, fun y hyV => (hfg ⟨y, hyV⟩).1, Subtype.eq <| funext fun y => by + obtain ⟨h1, (h2 : g • s.1 ⟨y, _⟩ = LocalizedModule.mk f 1)⟩ := hfg y + exact show LocalizedModule.mk f ⟨g, by exact h1⟩ = s.1 (iVU y) by + set x := s.1 (iVU y); change g • x = _ at h2; clear_value x + induction x using LocalizedModule.induction_on with + | h a b => + rw [LocalizedModule.smul'_mk, LocalizedModule.mk_eq] at h2 + obtain ⟨c, hc⟩ := h2 + exact LocalizedModule.mk_eq.mpr ⟨c, by simpa using hc.symm⟩⟩ + +@[simp] +theorem res_const (f : M) (g : R) (U hu V hv i) : + (tildeInModuleCat M).map i (const M f g U hu) = const M f g V hv := + rfl + +@[simp] +theorem localizationToStalk_mk (x : PrimeSpectrum.Top R) (f : M) (s : x.asIdeal.primeCompl) : + localizationToStalk M x (LocalizedModule.mk f s) = + (tildeInModuleCat M).germ (PrimeSpectrum.basicOpen (s : R)) x s.2 + (const M f s (PrimeSpectrum.basicOpen s) fun _ => id) := + (Module.End_isUnit_iff _ |>.1 (isUnit_toStalk M x s)).injective <| by + erw [← LinearMap.mul_apply] + simp only [IsUnit.mul_val_inv, LinearMap.one_apply, Module.algebraMap_end_apply] + show (M.tildeInModuleCat.germ ⊤ x ⟨⟩) ((toOpen M ⊤) f) = _ + rw [← map_smul] + fapply TopCat.Presheaf.germ_ext (W := PrimeSpectrum.basicOpen s.1) (hxW := s.2) + · exact homOfLE le_top + · exact 𝟙 _ + refine Subtype.eq <| funext fun y => show LocalizedModule.mk f 1 = _ from ?_ + show LocalizedModule.mk f 1 = s.1 • LocalizedModule.mk f _ + rw [LocalizedModule.smul'_mk, LocalizedModule.mk_eq] + exact ⟨1, by simp⟩ + +/-- +The isomorphism of `R`-modules from the stalk of `M^~` at `x` to the localization of `M` at the +prime ideal corresponding to `x`. +-/ +@[simps] +noncomputable def stalkIso (x : PrimeSpectrum.Top R) : + TopCat.Presheaf.stalk (tildeInModuleCat M) x ≅ + ModuleCat.of R (LocalizedModule x.asIdeal.primeCompl M) where + hom := stalkToFiberLinearMap M x + inv := localizationToStalk M x + hom_inv_id := TopCat.Presheaf.stalk_hom_ext _ fun U hxU ↦ ext _ fun s ↦ by + show localizationToStalk M x (stalkToFiberLinearMap M x (M.tildeInModuleCat.germ U x hxU s)) = + M.tildeInModuleCat.germ U x hxU s + rw [stalkToFiberLinearMap_germ] + obtain ⟨V, hxV, iVU, f, g, (hg : V ≤ PrimeSpectrum.basicOpen _), hs⟩ := + exists_const _ _ s x hxU + rw [← res_apply M U V iVU s ⟨x, hxV⟩, ← hs, const_apply, localizationToStalk_mk] + exact (tildeInModuleCat M).germ_ext V hxV (homOfLE hg) iVU <| hs ▸ rfl + inv_hom_id := by ext x; exact x.induction_on (fun _ _ => by simp) + +instance (x : PrimeSpectrum.Top R) : + IsLocalizedModule x.asIdeal.primeCompl (toStalk M x) := by + convert IsLocalizedModule.of_linearEquiv + (hf := localizedModuleIsLocalizedModule (M := M) x.asIdeal.primeCompl) + (e := (stalkIso M x).symm.toLinearEquiv) + simp only [of_coe, show (stalkIso M x).symm.toLinearEquiv.toLinearMap = (stalkIso M x).inv by rfl, + stalkIso_inv] + erw [LocalizedModule.lift_comp] + end Tilde end ModuleCat From 721c01e4de26b2bb91e0c32ddcdcdf842e97c405 Mon Sep 17 00:00:00 2001 From: FR Date: Sat, 19 Oct 2024 19:27:05 +0000 Subject: [PATCH 190/505] chore(Data/Nat/{Bit, Bitwise}): make `Nat.bit_mod_two` `Nat.bit_eq_zero_iff` be `simp` (#17924) Also deprecate `Nat.bit_eq_zero`. It is identical to `Nat.bit_eq_zero_iff`. --- Mathlib/Data/Nat/Bits.lean | 1 + Mathlib/Data/Nat/Bitwise.lean | 13 +++++-------- Mathlib/Data/Nat/Multiplicity.lean | 7 +++---- 3 files changed, 9 insertions(+), 12 deletions(-) diff --git a/Mathlib/Data/Nat/Bits.lean b/Mathlib/Data/Nat/Bits.lean index 71f01660fc34f..6ea51b387201e 100644 --- a/Mathlib/Data/Nat/Bits.lean +++ b/Mathlib/Data/Nat/Bits.lean @@ -328,6 +328,7 @@ theorem bit_cases_on_inj {motive : ℕ → Sort u} (H₁ H₂ : ∀ b n, motive ((fun n => bitCasesOn n H₁) = fun n => bitCasesOn n H₂) ↔ H₁ = H₂ := bit_cases_on_injective.eq_iff +@[simp] theorem bit_eq_zero_iff {n : ℕ} {b : Bool} : bit b n = 0 ↔ n = 0 ∧ b = false := by constructor · cases b <;> simp [Nat.bit]; omega diff --git a/Mathlib/Data/Nat/Bitwise.lean b/Mathlib/Data/Nat/Bitwise.lean index cce15c0d55eee..07873ec6fd976 100644 --- a/Mathlib/Data/Nat/Bitwise.lean +++ b/Mathlib/Data/Nat/Bitwise.lean @@ -88,19 +88,18 @@ lemma bitwise_bit {f : Bool → Bool → Bool} (h : f false false = false := by cases a <;> cases b <;> simp [h2, h4] <;> split_ifs <;> simp_all (config := {decide := true}) [two_mul] +@[simp] lemma bit_mod_two (a : Bool) (x : ℕ) : bit a x % 2 = a.toNat := by cases a <;> simp [bit_val, mul_add_mod] -@[simp] lemma bit_mod_two_eq_zero_iff (a x) : bit a x % 2 = 0 ↔ !a := by - simp [bit_mod_two] + simp -@[simp] lemma bit_mod_two_eq_one_iff (a x) : bit a x % 2 = 1 ↔ a := by - simp [bit_mod_two] + simp @[simp] theorem lor_bit : ∀ a m b n, bit a m ||| bit b n = bit (a || b) (m ||| n) := @@ -142,12 +141,10 @@ theorem bit_false : bit false = (2 * ·) := theorem bit_true : bit true = (2 * · + 1) := rfl -@[simp] -theorem bit_eq_zero {n : ℕ} {b : Bool} : n.bit b = 0 ↔ n = 0 ∧ b = false := by - cases b <;> simp [bit, Nat.mul_eq_zero] +@[deprecated (since := "2024-10-19")] alias bit_eq_zero := bit_eq_zero_iff theorem bit_ne_zero_iff {n : ℕ} {b : Bool} : n.bit b ≠ 0 ↔ n = 0 → b = true := by - simpa only [not_and, Bool.not_eq_false] using (@bit_eq_zero n b).not + simp /-- An alternative for `bitwise_bit` which replaces the `f false false = false` assumption with assumptions that neither `bit a m` nor `bit b n` are `0` diff --git a/Mathlib/Data/Nat/Multiplicity.lean b/Mathlib/Data/Nat/Multiplicity.lean index e86d693f824f5..61b86c32c88a8 100644 --- a/Mathlib/Data/Nat/Multiplicity.lean +++ b/Mathlib/Data/Nat/Multiplicity.lean @@ -6,7 +6,6 @@ Authors: Chris Hughes import Mathlib.Algebra.BigOperators.Intervals import Mathlib.Algebra.GeomSum import Mathlib.Algebra.Order.Ring.Abs -import Mathlib.Data.Nat.Bitwise import Mathlib.Data.Nat.Log import Mathlib.Data.Nat.Prime.Defs import Mathlib.Data.Nat.Digits @@ -282,10 +281,10 @@ theorem emultiplicity_two_factorial_lt : ∀ {n : ℕ} (_ : n ≠ 0), emultiplic · intro b n ih h by_cases hn : n = 0 · subst hn - simp only [ne_eq, bit_eq_zero, true_and, Bool.not_eq_false] at h - simp only [h, bit_true, factorial, mul_one, Nat.isUnit_iff, cast_one] + simp only [ne_eq, bit_eq_zero_iff, true_and, Bool.not_eq_false] at h + simp only [h, factorial, mul_one, Nat.isUnit_iff, cast_one] rw [Prime.emultiplicity_one] - · simp [zero_lt_one] + · exact zero_lt_one · decide have : emultiplicity 2 (2 * n)! < (2 * n : ℕ) := by rw [prime_two.emultiplicity_factorial_mul] From fee48cc34cea5902fd163549acb01e413d7e7243 Mon Sep 17 00:00:00 2001 From: Christian Merten Date: Sat, 19 Oct 2024 19:57:03 +0000 Subject: [PATCH 191/505] =?UTF-8?q?feat(AlgebraicGeometry/ResidueField):?= =?UTF-8?q?=20classification=20of=20`Spec=20K=20=E2=9F=B6=20X`=20with=20`K?= =?UTF-8?q?`=20a=20field=20(#17768)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR adds more API for residue fields and in particular the classification of morphisms `Spec K ⟶ X` with `K` a field. From the valuative criterion project. Co-authored-by: Qi Ge Co-authored-by: Andrew Yang Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> Co-authored-by: Christian Merten --- .../Morphisms/ClosedImmersion.lean | 6 +- .../Morphisms/Preimmersion.lean | 27 +++++ Mathlib/AlgebraicGeometry/ResidueField.lean | 108 +++++++++++++++++- Mathlib/AlgebraicGeometry/Stalk.lean | 15 ++- 4 files changed, 150 insertions(+), 6 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean b/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean index 932325058e426..03d557fa0c469 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean @@ -3,12 +3,10 @@ Copyright (c) 2023 Jonas van der Schaaf. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Amelia Livingston, Christian Merten, Jonas van der Schaaf -/ -import Mathlib.AlgebraicGeometry.Morphisms.Preimmersion import Mathlib.AlgebraicGeometry.Morphisms.QuasiSeparated import Mathlib.AlgebraicGeometry.Morphisms.RingHomProperties import Mathlib.AlgebraicGeometry.Morphisms.FiniteType import Mathlib.AlgebraicGeometry.ResidueField -import Mathlib.RingTheory.RingHom.Surjective /-! @@ -134,6 +132,10 @@ theorem of_comp {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) [IsClosedImmersion simp_rw [Scheme.stalkMap_comp] at h exact Function.Surjective.of_comp h +instance Spec_map_residue {X : Scheme.{u}} (x) : IsClosedImmersion (Spec.map (X.residue x)) := + IsClosedImmersion.spec_of_surjective (X.residue x) + Ideal.Quotient.mk_surjective + instance {X Y : Scheme} (f : X ⟶ Y) [IsClosedImmersion f] : QuasiCompact f where isCompact_preimage _ _ hU' := base_closed.isCompact_preimage hU' diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Preimmersion.lean b/Mathlib/AlgebraicGeometry/Morphisms/Preimmersion.lean index 9336dbb858cde..f504cc230196f 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Preimmersion.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Preimmersion.lean @@ -5,6 +5,7 @@ Authors: Andrew Yang -/ import Mathlib.AlgebraicGeometry.Morphisms.UnderlyingMap import Mathlib.RingTheory.RingHom.Surjective +import Mathlib.RingTheory.SurjectiveOnStalks /-! @@ -95,6 +96,32 @@ theorem comp_iff {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) [IsPreimmersion g] IsPreimmersion (f ≫ g) ↔ IsPreimmersion f := ⟨fun _ ↦ of_comp f g, fun _ ↦ inferInstance⟩ +lemma Spec_map_iff {R S : CommRingCat.{u}} (f : R ⟶ S) : + IsPreimmersion (Spec.map f) ↔ Embedding (PrimeSpectrum.comap f) ∧ f.SurjectiveOnStalks := by + haveI : (RingHom.toMorphismProperty <| fun f ↦ Function.Surjective f).RespectsIso := by + rw [← RingHom.toMorphismProperty_respectsIso_iff] + exact RingHom.surjective_respectsIso + refine ⟨fun ⟨h₁, h₂⟩ ↦ ⟨h₁, ?_⟩, fun ⟨h₁, h₂⟩ ↦ ⟨h₁, fun (x : PrimeSpectrum S) ↦ ?_⟩⟩ + · intro p hp + let e := Scheme.arrowStalkMapSpecIso f ⟨p, hp⟩ + apply ((RingHom.toMorphismProperty <| fun f ↦ Function.Surjective f).arrow_mk_iso_iff e).mp + exact h₂ ⟨p, hp⟩ + · let e := Scheme.arrowStalkMapSpecIso f x + apply ((RingHom.toMorphismProperty <| fun f ↦ Function.Surjective f).arrow_mk_iso_iff e).mpr + exact h₂ x.asIdeal x.isPrime + +lemma mk_Spec_map {R S : CommRingCat.{u}} {f : R ⟶ S} + (h₁ : Embedding (PrimeSpectrum.comap f)) (h₂ : f.SurjectiveOnStalks) : + IsPreimmersion (Spec.map f) := + (Spec_map_iff f).mpr ⟨h₁, h₂⟩ + +lemma of_isLocalization {R S : Type u} [CommRing R] (M : Submonoid R) [CommRing S] + [Algebra R S] [IsLocalization M S] : + IsPreimmersion (Spec.map (CommRingCat.ofHom <| algebraMap R S)) := + IsPreimmersion.mk_Spec_map + (PrimeSpectrum.localization_comap_embedding (R := R) S M) + (RingHom.surjectiveOnStalks_of_isLocalization (M := M) S) + end IsPreimmersion end AlgebraicGeometry diff --git a/Mathlib/AlgebraicGeometry/ResidueField.lean b/Mathlib/AlgebraicGeometry/ResidueField.lean index ac1d5d1e8dcc2..c7fcb54274fd0 100644 --- a/Mathlib/AlgebraicGeometry/ResidueField.lean +++ b/Mathlib/AlgebraicGeometry/ResidueField.lean @@ -3,8 +3,8 @@ Copyright (c) 2024 Andrew Yang. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang -/ -import Mathlib.Geometry.RingedSpace.LocallyRingedSpace.ResidueField import Mathlib.AlgebraicGeometry.Stalk +import Mathlib.Geometry.RingedSpace.LocallyRingedSpace.ResidueField /-! @@ -20,12 +20,15 @@ The following are in the `AlgebraicGeometry.Scheme` namespace: - `AlgebraicGeometry.Scheme.Hom.residueFieldMap`: A morphism of schemes induce a homomorphism of residue fields. - `AlgebraicGeometry.Scheme.fromSpecResidueField`: The canonical map `Spec κ(x) ⟶ X`. +- `AlgebraicGeometry.Scheme.SpecToEquivOfField`: morphisms `Spec K ⟶ X` for a field `K` correspond + to pairs of `x : X` with embedding `κ(x) ⟶ K`. + -/ universe u -open CategoryTheory TopologicalSpace Opposite +open CategoryTheory TopologicalSpace Opposite LocalRing noncomputable section @@ -45,12 +48,37 @@ instance (x : X) : Field (X.residueField x) := def residue (X : Scheme.{u}) (x) : X.presheaf.stalk x ⟶ X.residueField x := LocalRing.residue _ +/-- See `AlgebraicGeometry.IsClosedImmersion.Spec_map_residue` for the stronger result that +`Spec.map (X.residue x)` is a closed immersion. -/ +instance {X : Scheme.{u}} (x) : IsPreimmersion (Spec.map (X.residue x)) := + IsPreimmersion.mk_Spec_map + (PrimeSpectrum.closedEmbedding_comap_of_surjective _ _ (Ideal.Quotient.mk_surjective)).embedding + (RingHom.surjectiveOnStalks_of_surjective (Ideal.Quotient.mk_surjective)) + +@[simp] +lemma Spec_map_residue_apply {X : Scheme.{u}} (x : X) (s : Spec (X.residueField x)) : + (Spec.map (X.residue x)).base s = closedPoint (X.presheaf.stalk x) := + LocalRing.PrimeSpectrum.comap_residue _ s + lemma residue_surjective (X : Scheme.{u}) (x) : Function.Surjective (X.residue x) := Ideal.Quotient.mk_surjective instance (X : Scheme.{u}) (x) : Epi (X.residue x) := ConcreteCategory.epi_of_surjective _ (X.residue_surjective x) +/-- If `K` is a field and `f : 𝒪_{X, x} ⟶ K` is a ring map, then this is the induced +map `κ(x) ⟶ K`. -/ +def descResidueField {K : Type u} [Field K] {X : Scheme.{u}} {x : X} + (f : X.presheaf.stalk x ⟶ .of K) [IsLocalHom f] : + X.residueField x ⟶ .of K := + LocalRing.ResidueField.lift (S := K) f + +@[reassoc (attr := simp)] +lemma residue_descResidueField {K : Type u} [Field K] {X : Scheme.{u}} {x} + (f : X.presheaf.stalk x ⟶ .of K) [IsLocalHom f] : + X.residue x ≫ X.descResidueField f = f := + RingHom.ext fun _ ↦ rfl + /-- If `U` is an open of `X` containing `x`, we have a canonical ring map from the sections over `U` to the residue field of `x`. @@ -171,6 +199,10 @@ lemma residue_residueFieldCongr (X : Scheme) {x y : X} (h : x = y) : subst h simp +lemma Hom.residueFieldMap_congr {f g : X ⟶ Y} (e : f = g) (x : X) : + f.residueFieldMap x = (Y.residueFieldCongr (by subst e; rfl)).hom ≫ g.residueFieldMap x := by + subst e; simp + end congr section fromResidueField @@ -178,7 +210,12 @@ section fromResidueField /-- The canonical map `Spec κ(x) ⟶ X`. -/ def fromSpecResidueField (X : Scheme) (x : X) : Spec (X.residueField x) ⟶ X := - Spec.map (CommRingCat.ofHom (X.residue x)) ≫ X.fromSpecStalk x + Spec.map (X.residue x) ≫ X.fromSpecStalk x + +instance {X : Scheme.{u}} (x : X) : IsPreimmersion (X.fromSpecResidueField x) := by + dsimp only [Scheme.fromSpecResidueField] + rw [IsPreimmersion.comp_iff] + infer_instance @[reassoc (attr := simp)] lemma residueFieldCongr_fromSpecResidueField {x y : X} (h : x = y) : @@ -186,8 +223,73 @@ lemma residueFieldCongr_fromSpecResidueField {x y : X} (h : x = y) : X.fromSpecResidueField _ := by subst h; simp +@[reassoc] +lemma Hom.Spec_map_residueFieldMap_fromSpecResidueField (x : X) : + Spec.map (f.residueFieldMap x) ≫ Y.fromSpecResidueField _ = + X.fromSpecResidueField x ≫ f := by + dsimp only [fromSpecResidueField] + rw [Category.assoc, ← Spec_map_stalkMap_fromSpecStalk, ← Spec.map_comp_assoc, + ← Spec.map_comp_assoc] + rfl + +@[simp] +lemma fromSpecResidueField_apply (x : X.carrier) (s : Spec (X.residueField x)) : + (X.fromSpecResidueField x).base s = x := by + simp [fromSpecResidueField] + +lemma range_fromSpecResidueField (x : X.carrier) : + Set.range (X.fromSpecResidueField x).base = {x} := by + ext s + simp only [Set.mem_range, fromSpecResidueField_apply, Set.mem_singleton_iff, eq_comm (a := s)] + constructor + · rintro ⟨-, h⟩ + exact h + · rintro rfl + exact ⟨closedPoint (X.residueField x), rfl⟩ + +lemma descResidueField_fromSpecResidueField {K : Type*} [Field K] (X : Scheme) {x} + (f : X.presheaf.stalk x ⟶ .of K) [IsLocalHom f] : + Spec.map (X.descResidueField f) ≫ + X.fromSpecResidueField x = Spec.map f ≫ X.fromSpecStalk x := by + simp [fromSpecResidueField, ← Spec.map_comp_assoc] + +lemma descResidueField_stalkClosedPointTo_fromSpecResidueField + (K : Type u) [Field K] (X : Scheme.{u}) (f : Spec (.of K) ⟶ X) : + Spec.map (X.descResidueField (Scheme.stalkClosedPointTo f)) ≫ + X.fromSpecResidueField (f.base (closedPoint K)) = f := by + erw [X.descResidueField_fromSpecResidueField] + rw [Scheme.Spec_stalkClosedPointTo_fromSpecStalk] + end fromResidueField +/-- A helper lemma to work with `AlgebraicGeometry.Scheme.SpecToEquivOfField`. -/ +lemma SpecToEquivOfField_eq_iff {K : Type*} [Field K] {X : Scheme} + {f₁ f₂ : Σ x : X.carrier, X.residueField x ⟶ .of K} : + f₁ = f₂ ↔ ∃ e : f₁.1 = f₂.1, f₁.2 = (X.residueFieldCongr e).hom ≫ f₂.2 := by + constructor + · rintro rfl + simp + · obtain ⟨f, _⟩ := f₁ + obtain ⟨g, _⟩ := f₂ + rintro ⟨(rfl : f = g), h⟩ + simpa + +/-- For a field `K` and a scheme `X`, the morphisms `Spec K ⟶ X` bijectively correspond +to pairs of points `x` of `X` and embeddings `κ(x) ⟶ K`. -/ +def SpecToEquivOfField (K : Type u) [Field K] (X : Scheme.{u}) : + (Spec (.of K) ⟶ X) ≃ Σ x, X.residueField x ⟶ .of K where + toFun f := + ⟨_, X.descResidueField (Scheme.stalkClosedPointTo f)⟩ + invFun xf := Spec.map xf.2 ≫ X.fromSpecResidueField xf.1 + left_inv := Scheme.descResidueField_stalkClosedPointTo_fromSpecResidueField K X + right_inv f := by + rw [SpecToEquivOfField_eq_iff] + simp only [CommRingCat.coe_of, Scheme.comp_coeBase, TopCat.coe_comp, Function.comp_apply, + Scheme.fromSpecResidueField_apply, exists_true_left] + rw [← Spec.map_inj, Spec.map_comp, ← cancel_mono (X.fromSpecResidueField _)] + erw [Scheme.descResidueField_stalkClosedPointTo_fromSpecResidueField] + simp + end Scheme end AlgebraicGeometry diff --git a/Mathlib/AlgebraicGeometry/Stalk.lean b/Mathlib/AlgebraicGeometry/Stalk.lean index 251d379d23a14..92b56c3849cc3 100644 --- a/Mathlib/AlgebraicGeometry/Stalk.lean +++ b/Mathlib/AlgebraicGeometry/Stalk.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang, Fangming Li -/ import Mathlib.AlgebraicGeometry.AffineScheme +import Mathlib.AlgebraicGeometry.Morphisms.Preimmersion /-! # Stalks of a Scheme @@ -68,6 +69,19 @@ noncomputable def Scheme.fromSpecStalk (X : Scheme) (x : X) : theorem IsAffineOpen.fromSpecStalk_eq_fromSpecStalk {x : X} (hxU : x ∈ U) : hU.fromSpecStalk hxU = X.fromSpecStalk x := fromSpecStalk_eq .. +instance IsAffineOpen.fromSpecStalk_isPreimmersion {X : Scheme.{u}} {U : Opens X} + (hU : IsAffineOpen U) (x : X) (hx : x ∈ U) : IsPreimmersion (hU.fromSpecStalk hx) := by + dsimp [IsAffineOpen.fromSpecStalk] + haveI : IsPreimmersion (Spec.map (X.presheaf.germ U x hx)) := + letI : Algebra Γ(X, U) (X.presheaf.stalk x) := (X.presheaf.germ U x hx).toAlgebra + haveI := hU.isLocalization_stalk ⟨x, hx⟩ + IsPreimmersion.of_isLocalization (R := Γ(X, U)) (S := X.presheaf.stalk x) + (hU.primeIdealOf ⟨x, hx⟩).asIdeal.primeCompl + apply IsPreimmersion.comp + +instance {X : Scheme.{u}} (x : X) : IsPreimmersion (X.fromSpecStalk x) := + IsAffineOpen.fromSpecStalk_isPreimmersion _ _ _ + lemma IsAffineOpen.fromSpecStalk_closedPoint {U : Opens X} (hU : IsAffineOpen U) {x : X} (hxU : x ∈ U) : (hU.fromSpecStalk hxU).base (closedPoint (X.presheaf.stalk x)) = x := by @@ -184,7 +198,6 @@ section stalkClosedPointTo variable {R} (f : Spec R ⟶ X) - namespace Scheme /-- From 227950d6190a304c6d697f37d74425bebe12917d Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Sat, 19 Oct 2024 21:32:09 +0000 Subject: [PATCH 192/505] feat: topology on module over topological ring (#16895) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Given a module over a topological ring, we define the module topology on this module to be the finest module making it into a topological module (i.e. the finest topology making addition and scalar multiplication continuous). NB this PR gave rise to a discussion about whether `⨆ a ∈ s, f a` or `sSup (f '' s)` should be the simp normal form, which was discussed on [Zulip](https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/normal.20form.20for.20pullback.20of.20Inf.20of.20topologies/near/472676441) without any clear conclusion. Co-authored-by: Eric Wieser --- Mathlib.lean | 1 + Mathlib/Topology/Algebra/Module/Basic.lean | 13 + .../Algebra/Module/ModuleTopology.lean | 265 ++++++++++++++++++ Mathlib/Topology/Order.lean | 10 + 4 files changed, 289 insertions(+) create mode 100644 Mathlib/Topology/Algebra/Module/ModuleTopology.lean diff --git a/Mathlib.lean b/Mathlib.lean index 5fcbb05388323..c75bbf3dd05cd 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4553,6 +4553,7 @@ import Mathlib.Topology.Algebra.Module.Determinant import Mathlib.Topology.Algebra.Module.FiniteDimension import Mathlib.Topology.Algebra.Module.LinearPMap import Mathlib.Topology.Algebra.Module.LocallyConvex +import Mathlib.Topology.Algebra.Module.ModuleTopology import Mathlib.Topology.Algebra.Module.Multilinear.Basic import Mathlib.Topology.Algebra.Module.Multilinear.Bounded import Mathlib.Topology.Algebra.Module.Multilinear.Topology diff --git a/Mathlib/Topology/Algebra/Module/Basic.lean b/Mathlib/Topology/Algebra/Module/Basic.lean index 4296c9b917999..943df043c2c09 100644 --- a/Mathlib/Topology/Algebra/Module/Basic.lean +++ b/Mathlib/Topology/Algebra/Module/Basic.lean @@ -12,6 +12,7 @@ import Mathlib.Algebra.Algebra.Defs import Mathlib.LinearAlgebra.Projection import Mathlib.LinearAlgebra.Pi import Mathlib.LinearAlgebra.Finsupp +import Mathlib.Algebra.Module.Opposites /-! # Theory of topological modules and continuous linear maps. @@ -2379,4 +2380,16 @@ end Submodule end Quotient +namespace MulOpposite + +variable (R : Type*) [Semiring R] [τR : TopologicalSpace R] [TopologicalSemiring R] + {M : Type*} [AddCommMonoid M] [Module R M] [TopologicalSpace M] [ContinuousSMul R M] + +/-- The function `op` is a continuous linear equivalence. -/ +@[simps!] +def opContinuousLinearEquiv : M ≃L[R] Mᵐᵒᵖ where + __ := MulOpposite.opLinearEquiv R + +end MulOpposite + set_option linter.style.longFile 2500 diff --git a/Mathlib/Topology/Algebra/Module/ModuleTopology.lean b/Mathlib/Topology/Algebra/Module/ModuleTopology.lean new file mode 100644 index 0000000000000..6fdfe9edc0b90 --- /dev/null +++ b/Mathlib/Topology/Algebra/Module/ModuleTopology.lean @@ -0,0 +1,265 @@ +/- +Copyright (c) 2024 Kevin Buzzard. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kevin Buzzard, Will Sawin +-/ +import Mathlib.Topology.Algebra.Module.Basic + +/-! +# A "module topology" for modules over a topological ring + +If `R` is a topological ring acting on an additive abelian group `A`, we define the +*module topology* to be the finest topology on `A` making both the maps +`• : R × A → A` and `+ : A × A → A` continuous (with all the products having the +product topology). Note that `- : A → A` is also automatically continuous as it is `a ↦ (-1) • a`. + +This topology was suggested by Will Sawin [here](https://mathoverflow.net/a/477763/1384). + + +## Mathematical details + +I (buzzard) don't know of any reference for this other than Sawin's mathoverflow answer, +so I expand some of the details here. + +First note that the definition makes sense in far more generality (for example `R` just needs to +be a topological space acting on an additive monoid). + +Next note that there *is* a finest topology with this property! Indeed, topologies on a fixed +type form a complete lattice (infinite infs and sups exist). So if `τ` is the Inf of all +the topologies on `A` which make `+` and `•` continuous, then the claim is that `+` and `•` +are still continuous for `τ` (note that topologies are ordered so that finer topologies +are smaller). To show `+ : A × A → A` is continuous we equivalently need to show +that the pushforward of the product topology `τ × τ` along `+` is `≤ τ`, and because `τ` is +the greatest lower bound of the topologies making `•` and `+` continuous, +it suffices to show that it's `≤ σ` for any topology `σ` on `A` which makes `+` and `•` continuous. +However pushforward and products are monotone, so `τ × τ ≤ σ × σ`, and the pushforward of +`σ × σ` is `≤ σ` because that's precisely the statement that `+` is continuous for `σ`. +The proof for `•` follows mutatis mutandis. + +A *topological module* for a topological ring `R` is an `R`-module `A` with a topology +making `+` and `•` continuous. The discussion so far has shown that the module topology makes +an `R`-module `A` into a topological module, and moreover is the finest topology with this property. + +A crucial observation is that if `M` is a topological `R`-module, if `A` is an `R`-module with no +topology, and if `φ : A → M` is linear, then the pullback of `M`'s topology to `A` is a topology +making `A` into a topological module. Let's for example check that `•` is continuous. +If `U ⊆ A` is open then by definition of the pullback topology, `U = φ⁻¹(V)` for some open `V ⊆ M`, +and now the pullback of `U` under `•` is just the pullback along the continuous map +`id × φ : R × A → R × M` of the preimage of `V` under the continuous map `• : R × M → M`, +so it's open. The proof for `+` is similar. + +As a consequence of this, we see that if `φ : A → M` is a linear map between topological `R`-modules +modules and if `A` has the module topology, then `φ` is automatically continuous. +Indeed the argument above shows that if `A → M` is linear then the module +topology on `A` is `≤` the pullback of the module topology on `M` (because it's the inf of a set +containing this topology) which is the definition of continuity. + +We also deduce that the module topology is a functor from the category of `R`-modules +(`R` a topological ring) to the category of topological `R`-modules, and it is perhaps +unsurprising that this is an adjoint to the forgetful functor. Indeed, if `A` is an `R`-module +and `M` is a topological `R`-module, then the previous paragraph shows that +the linear maps `A → M` are precisely the continuous linear maps +from (`A` with its module topology) to `M`, so the module topology is a left adjoint +to the forgetful functor. + +This file develops the theory of the module topology. We prove that the module topology on +`R` as a module over itself is `R`'s original topology, and that the module topology +is isomorphism-invariant. + +## Main theorems + +* `TopologicalSemiring.toIsModuleTopology : IsModuleTopology R R`. The module + topology on `R` is `R`'s topology. +* `IsModuleTopology.iso [IsModuleTopology R A] (e : A ≃L[R] B) : IsModuleTopology R B`. If `A` and + `B` are `R`-modules with topologies, if `e` is a topological isomorphism between them, + and if `A` has the module topology, then `B` has the module topology. + +## TODO + +Forthcoming PRs from the FLT repo will show that the module topology on a (binary or finite) product +of modules is the product of the module topologies, and that the module topology on the quotient +of a module `M` is the quotient topology when `M` is equipped with the module topology. + +We will also show the slightly more subtle result that if `M`, `N` and `P` are `R`-modules +equipped with the module topology and if furthermore `M` is finite as an `R`-module, +then any bilinear map `M × N → P` is continuous. + +As a consequence of this, we deduce that if `R` is a commutative topological ring +and `A` is an `R`-algebra of finite type as `R`-module, then `A` with its module +topology becomes a topological ring (i.e., multiplication is continuous). + +Other TODOs (not done in the FLT repo): + +* The module topology is a functor from the category of `R`-modules + to the category of topological `R`-modules, and it's an adjoint to the forgetful functor. + +-/ + +section basics + +/- +This section is just boilerplate, defining the module topology and making +some infrastructure. Note that we don't even need that `R` is a ring +in the main definitions, just that it acts on `A`. +-/ +variable (R : Type*) [TopologicalSpace R] (A : Type*) [Add A] [SMul R A] + +/-- The module topology, for a module `A` over a topological ring `R`. It's the finest topology +making addition and the `R`-action continuous, or equivalently the finest topology making `A` +into a topological `R`-module. More precisely it's the Inf of the set of +topologies with these properties; theorems `continuousSMul` and `continuousAdd` show +that the module topology also has these properties. -/ +abbrev moduleTopology : TopologicalSpace A := + sInf {t | @ContinuousSMul R A _ _ t ∧ @ContinuousAdd A t _} + +/-- A class asserting that the topology on a module over a topological ring `R` is +the module topology. See `moduleTopology` for more discussion of the module topology. -/ +class IsModuleTopology [τA : TopologicalSpace A] : Prop where + /-- Note that this should not be used directly, and `eq_moduleTopology`, which takes `R` and `A` + explicitly, should be used instead. -/ + eq_moduleTopology' : τA = moduleTopology R A + +theorem eq_moduleTopology [τA : TopologicalSpace A] [IsModuleTopology R A] : + τA = moduleTopology R A := + IsModuleTopology.eq_moduleTopology' (R := R) (A := A) + +/-- Scalar multiplication `• : R × A → A` is continuous if `R` is a topological +ring, and `A` is an `R` module with the module topology. -/ +theorem ModuleTopology.continuousSMul : @ContinuousSMul R A _ _ (moduleTopology R A) := + /- Proof: We need to prove that the product topology is finer than the pullback + of the module topology. But the module topology is an Inf and thus a limit, + and pullback is a right adjoint, so it preserves limits. + We must thus show that the product topology is finer than an Inf, so it suffices + to show it's a lower bound, which is not hard. All this is wrapped into + `continuousSMul_sInf`. + -/ + continuousSMul_sInf fun _ h ↦ h.1 + +/-- Addition `+ : A × A → A` is continuous if `R` is a topological +ring, and `A` is an `R` module with the module topology. -/ +theorem ModuleTopology.continuousAdd : @ContinuousAdd A (moduleTopology R A) _ := + continuousAdd_sInf fun _ h ↦ h.2 + +instance IsModuleTopology.toContinuousSMul [TopologicalSpace A] [IsModuleTopology R A] : + ContinuousSMul R A := eq_moduleTopology R A ▸ ModuleTopology.continuousSMul R A + +-- this can't be an instance because typclass inference can't be expected to find `R`. +theorem IsModuleTopology.toContinuousAdd [TopologicalSpace A] [IsModuleTopology R A] : + ContinuousAdd A := eq_moduleTopology R A ▸ ModuleTopology.continuousAdd R A + +/-- The module topology is `≤` any topology making `A` into a topological module. -/ +theorem moduleTopology_le [τA : TopologicalSpace A] [ContinuousSMul R A] [ContinuousAdd A] : + moduleTopology R A ≤ τA := sInf_le ⟨inferInstance, inferInstance⟩ + +end basics + +namespace IsModuleTopology + +section basics + +variable {R : Type*} [TopologicalSpace R] + {A : Type*} [Add A] [SMul R A] [τA : TopologicalSpace A] + +/-- If `A` is a topological `R`-module and the identity map from (`A` with its given +topology) to (`A` with the module topology) is continuous, then the topology on `A` is +the module topology. -/ +theorem of_continuous_id [ContinuousAdd A] [ContinuousSMul R A] + (h : @Continuous A A τA (moduleTopology R A) id): + IsModuleTopology R A where + -- The topologies are equal because each is finer than the other. One inclusion + -- follows from the continuity hypothesis; the other is because the module topology + -- is the inf of all the topologies making `A` a topological module. + eq_moduleTopology' := le_antisymm (continuous_id_iff_le.1 h) (moduleTopology_le _ _) + +/-- The zero module has the module topology. -/ +instance instSubsingleton [Subsingleton A] : IsModuleTopology R A where + eq_moduleTopology' := Subsingleton.elim _ _ + +end basics + +section iso + +variable {R : Type*} [τR : TopologicalSpace R] [Semiring R] +variable {A : Type*} [AddCommMonoid A] [Module R A] [τA : TopologicalSpace A] [IsModuleTopology R A] +variable {B : Type*} [AddCommMonoid B] [Module R B] [τB : TopologicalSpace B] + +/-- If `A` and `B` are `R`-modules, homeomorphic via an `R`-linear homeomorphism, and if +`A` has the module topology, then so does `B`. -/ +theorem iso (e : A ≃L[R] B) : IsModuleTopology R B where + eq_moduleTopology' := by + -- get these in before I start putting new topologies on A and B and have to use `@` + let g : A →ₗ[R] B := e + let g' : B →ₗ[R] A := e.symm + let h : A →+ B := e + let h' : B →+ A := e.symm + simp_rw [e.toHomeomorph.symm.inducing.1, eq_moduleTopology R A, moduleTopology, induced_sInf] + apply congr_arg + ext τ -- from this point on the definitions of `g`, `g'` etc above don't work without `@`. + rw [Set.mem_image] + constructor + · rintro ⟨σ, ⟨hσ1, hσ2⟩, rfl⟩ + exact ⟨continuousSMul_induced g', continuousAdd_induced h'⟩ + · rintro ⟨h1, h2⟩ + use τ.induced e + rw [induced_compose] + refine ⟨⟨continuousSMul_induced g, continuousAdd_induced h⟩, ?_⟩ + nth_rw 2 [← induced_id (t := τ)] + simp + +end iso + +section self + +variable (R : Type*) [Semiring R] [τR : TopologicalSpace R] [TopologicalSemiring R] + +/-! +We now fix once and for all a topological semiring `R`. + +We first prove that the module topology on `R` considered as a module over itself, +is `R`'s topology. +-/ + +/-- The topology on a topological semiring `R` agrees with the module topology when considering +`R` as an `R`-module in the obvious way (i.e., via `Semiring.toModule`). -/ +instance _root_.TopologicalSemiring.toIsModuleTopology : IsModuleTopology R R := by + /- By a previous lemma it suffices to show that the identity from (R,usual) to + (R, module topology) is continuous. -/ + apply of_continuous_id + /- + The idea needed here is to rewrite the identity function as the composite of `r ↦ (r,1)` + from `R` to `R × R`, and multiplication `R × R → R`. + -/ + rw [show (id : R → R) = (fun rs ↦ rs.1 • rs.2) ∘ (fun r ↦ (r, 1)) by ext; simp] + /- + It thus suffices to show that each of these maps are continuous. For this claim to even make + sense, we need to topologise `R × R`. The trick is to do this by giving the first `R` the usual + topology `τR` and the second `R` the module topology. To do this we have to "fight mathlib" + a bit with `@`, because there is more than one topology on `R` here. + -/ + apply @Continuous.comp R (R × R) R τR (@instTopologicalSpaceProd R R τR (moduleTopology R R)) + (moduleTopology R R) + · /- + The map R × R → R is `•`, so by a fundamental property of the module topology, + this is continuous. -/ + exact @continuous_smul _ _ _ _ (moduleTopology R R) <| ModuleTopology.continuousSMul .. + · /- + The map `R → R × R` sending `r` to `(r,1)` is a map into a product, so it suffices to show + that each of the two factors is continuous. But the first is the identity function + on `(R, usual topology)` and the second is a constant function. -/ + exact @Continuous.prod_mk _ _ _ _ (moduleTopology R R) _ _ _ continuous_id <| + @continuous_const _ _ _ (moduleTopology R R) _ + +end self + +section MulOpposite + +variable (R : Type*) [Semiring R] [τR : TopologicalSpace R] [TopologicalSemiring R] + +/-- The module topology coming from the action of the topological ring `Rᵐᵒᵖ` on `R` + (via `Semiring.toOppositeModule`, i.e. via `(op r) • m = m * r`) is `R`'s topology. -/ +instance _root_.TopologicalSemiring.toOppositeIsModuleTopology : IsModuleTopology Rᵐᵒᵖ R := + .iso (MulOpposite.opContinuousLinearEquiv Rᵐᵒᵖ).symm + +end MulOpposite + +end IsModuleTopology diff --git a/Mathlib/Topology/Order.lean b/Mathlib/Topology/Order.lean index 05a207e173994..173ade1245b1d 100644 --- a/Mathlib/Topology/Order.lean +++ b/Mathlib/Topology/Order.lean @@ -393,6 +393,11 @@ theorem induced_iInf {ι : Sort w} {t : ι → TopologicalSpace α} : (⨅ i, t i).induced g = ⨅ i, (t i).induced g := (gc_coinduced_induced g).u_iInf +@[simp] +theorem induced_sInf {s : Set (TopologicalSpace α)} : + TopologicalSpace.induced g (sInf s) = sInf (TopologicalSpace.induced g '' s) := by + rw [sInf_eq_iInf', sInf_image', induced_iInf] + @[simp] theorem coinduced_bot : (⊥ : TopologicalSpace α).coinduced f = ⊥ := (gc_coinduced_induced f).l_bot @@ -406,6 +411,11 @@ theorem coinduced_iSup {ι : Sort w} {t : ι → TopologicalSpace α} : (⨆ i, t i).coinduced f = ⨆ i, (t i).coinduced f := (gc_coinduced_induced f).l_iSup +@[simp] +theorem coinduced_sSup {s : Set (TopologicalSpace α)} : + TopologicalSpace.coinduced f (sSup s) = sSup ((TopologicalSpace.coinduced f) '' s) := by + rw [sSup_eq_iSup', sSup_image', coinduced_iSup] + theorem induced_id [t : TopologicalSpace α] : t.induced id = t := TopologicalSpace.ext <| funext fun s => propext <| ⟨fun ⟨_, hs, h⟩ => h ▸ hs, fun hs => ⟨s, hs, rfl⟩⟩ From e9f0a88e5333f1edc2843a9164fb035a96406f5e Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sat, 19 Oct 2024 21:53:49 +0000 Subject: [PATCH 193/505] feat(Normed/Group): add `nndist_one_right` etc (#17954) Add - `nndist_one_right`, `nndist_zero_right`, `nndist_one_left`, `nndist_zero_left`; - `edist_one_right`, `edist_zero_right`, `edist_one_left`, `edist_zero_left`. --- Mathlib/Analysis/Normed/Group/Basic.lean | 14 ++++++++++++++ Mathlib/MeasureTheory/Function/L1Space.lean | 6 ++---- 2 files changed, 16 insertions(+), 4 deletions(-) diff --git a/Mathlib/Analysis/Normed/Group/Basic.lean b/Mathlib/Analysis/Normed/Group/Basic.lean index e379a64547598..b3742ef73f777 100644 --- a/Mathlib/Analysis/Normed/Group/Basic.lean +++ b/Mathlib/Analysis/Normed/Group/Basic.lean @@ -635,6 +635,13 @@ theorem nndist_eq_nnnorm_div (a b : E) : nndist a b = ‖a / b‖₊ := alias nndist_eq_nnnorm := nndist_eq_nnnorm_sub +@[to_additive (attr := simp)] +theorem nndist_one_right (a : E) : nndist a 1 = ‖a‖₊ := by simp [nndist_eq_nnnorm_div] + +@[to_additive (attr := simp)] +theorem edist_one_right (a : E) : edist a 1 = ‖a‖₊ := by + rw [edist_nndist, nndist_one_right] + @[to_additive (attr := simp) nnnorm_zero] theorem nnnorm_one' : ‖(1 : E)‖₊ = 0 := NNReal.eq norm_one' @@ -653,6 +660,13 @@ theorem nnnorm_mul_le' (a b : E) : ‖a * b‖₊ ≤ ‖a‖₊ + ‖b‖₊ := theorem nnnorm_inv' (a : E) : ‖a⁻¹‖₊ = ‖a‖₊ := NNReal.eq <| norm_inv' a +@[to_additive (attr := simp)] +theorem nndist_one_left (a : E) : nndist 1 a = ‖a‖₊ := by simp [nndist_eq_nnnorm_div] + +@[to_additive (attr := simp)] +theorem edist_one_left (a : E) : edist 1 a = ‖a‖₊ := by + rw [edist_nndist, nndist_one_left] + open scoped symmDiff in @[to_additive] theorem nndist_mulIndicator (s t : Set α) (f : α → E) (x : α) : diff --git a/Mathlib/MeasureTheory/Function/L1Space.lean b/Mathlib/MeasureTheory/Function/L1Space.lean index 9994ae4d6e8b1..87c029742d760 100644 --- a/Mathlib/MeasureTheory/Function/L1Space.lean +++ b/Mathlib/MeasureTheory/Function/L1Space.lean @@ -1375,12 +1375,10 @@ theorem edist_toL1_toL1 (f g : α → β) (hf : Integrable f μ) (hg : Integrabl ENNReal.rpow_one, ne_eq, not_false_eq_true, div_self, ite_false] simp [edist_eq_coe_nnnorm_sub] -@[simp] theorem edist_toL1_zero (f : α → β) (hf : Integrable f μ) : edist (hf.toL1 f) 0 = ∫⁻ a, edist (f a) 0 ∂μ := by - simp only [toL1, Lp.edist_toLp_zero, eLpNorm, one_ne_zero, eLpNorm', one_toReal, ENNReal.rpow_one, - ne_eq, not_false_eq_true, div_self, ite_false] - simp [edist_eq_coe_nnnorm] + simp only [edist_zero_right, Lp.nnnorm_coe_ennreal, toL1_eq_mk, eLpNorm_aeeqFun] + apply eLpNorm_one_eq_lintegral_nnnorm variable {𝕜 : Type*} [NormedRing 𝕜] [Module 𝕜 β] [BoundedSMul 𝕜 β] From e0dccaab3f60adf16923fcc30de87ce08f296833 Mon Sep 17 00:00:00 2001 From: Thmoas-Guan <150537269+Thmoas-Guan@users.noreply.github.com> Date: Sat, 19 Oct 2024 22:21:59 +0000 Subject: [PATCH 194/505] chore(FieldTheory/Galois): Move infinite Galois (#17945) Split the original file `FieldTheory.Galois.Infinite` into `GaloisClosure` and `Profinite`, for spliting the basic constructions, profinite (stated categorically), and the fundamental theorem completely. Moves: - Mathlib.FieldTheory.Galois.Infinite -> Mathlib.FieldTheory.Galois.GaloisClosure - Mathlib.FieldTheory.Galois.Infinite -> Mathlib.FieldTheory.Galois.Profinite --- Mathlib.lean | 3 +- .../{Infinite.lean => GaloisClosure.lean} | 72 +--------------- Mathlib/FieldTheory/Galois/Profinite.lean | 84 +++++++++++++++++++ 3 files changed, 87 insertions(+), 72 deletions(-) rename Mathlib/FieldTheory/Galois/{Infinite.lean => GaloisClosure.lean} (63%) create mode 100644 Mathlib/FieldTheory/Galois/Profinite.lean diff --git a/Mathlib.lean b/Mathlib.lean index c75bbf3dd05cd..3ace00fab33be 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2730,7 +2730,8 @@ import Mathlib.FieldTheory.Finite.Trace import Mathlib.FieldTheory.Finiteness import Mathlib.FieldTheory.Fixed import Mathlib.FieldTheory.Galois.Basic -import Mathlib.FieldTheory.Galois.Infinite +import Mathlib.FieldTheory.Galois.GaloisClosure +import Mathlib.FieldTheory.Galois.Profinite import Mathlib.FieldTheory.IntermediateField.Algebraic import Mathlib.FieldTheory.IntermediateField.Basic import Mathlib.FieldTheory.IsAlgClosed.AlgebraicClosure diff --git a/Mathlib/FieldTheory/Galois/Infinite.lean b/Mathlib/FieldTheory/Galois/GaloisClosure.lean similarity index 63% rename from Mathlib/FieldTheory/Galois/Infinite.lean rename to Mathlib/FieldTheory/Galois/GaloisClosure.lean index d3c2345887e0d..c1f93236da4c0 100644 --- a/Mathlib/FieldTheory/Galois/Infinite.lean +++ b/Mathlib/FieldTheory/Galois/GaloisClosure.lean @@ -1,11 +1,9 @@ /- Copyright (c) 2024 Nailin Guan. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Nailin Guan, Yuyang Zhao, Jujian Zhang +Authors: Nailin Guan, Yuyang Zhao -/ -import Mathlib.Algebra.Category.Grp.FiniteGrp -import Mathlib.CategoryTheory.Category.Preorder import Mathlib.FieldTheory.NormalClosure import Mathlib.FieldTheory.SeparableClosure @@ -21,24 +19,12 @@ In a field extension `K/k` * `adjoin` : The finite Galois intermediate field obtained from the normal closure of adjoining a finite `s : Set K` to `k`. -* `finGaloisGroup L` : The (finite) Galois group `Gal(L/k)` associated to a - `L : FiniteGaloisIntermediateField k K` `L`. - -* `finGaloisGroupMap` : For `FiniteGaloisIntermediateField` s `L₁` and `L₂` with `L₂ ≤ L₁` - giving the restriction of `Gal(L₁/k)` to `Gal(L₂/k)` - -* `finGaloisGroupFunctor` : Mapping `FiniteGaloisIntermediateField` ordered by inverse inclusion to - its corresponding Galois Group as `FiniteGrp`. - # TODO * `FiniteGaloisIntermediateField` should be a `ConditionallyCompleteLattice` but isn't proved yet. -/ -open CategoryTheory Opposite -open scoped IntermediateField - variable (k K : Type*) [Field k] [Field K] [Algebra k K] /-- The type of intermediate fields of `K/k` that are finite and Galois over `k` -/ @@ -156,59 +142,3 @@ theorem adjoin_simple_map_algEquiv [IsGalois k K] (f : K ≃ₐ[k] K) (x : K) : adjoin_simple_map_algHom (f : K →ₐ[k] K) x end FiniteGaloisIntermediateField - -section Profinite - -variable {k K} - -/-- The (finite) Galois group `Gal(L / k)` associated to a -`L : FiniteGaloisIntermediateField k K` `L`. -/ -def FiniteGaloisIntermediateField.finGaloisGroup (L : FiniteGaloisIntermediateField k K) : - FiniteGrp := - letI := AlgEquiv.fintype k L - FiniteGrp.of <| L ≃ₐ[k] L - -/-- For `FiniteGaloisIntermediateField` s `L₁` and `L₂` with `L₂ ≤ L₁` - the restriction homomorphism from `Gal(L₁/k)` to `Gal(L₂/k)` -/ -noncomputable def finGaloisGroupMap {L₁ L₂ : (FiniteGaloisIntermediateField k K)ᵒᵖ} - (le : L₁ ⟶ L₂) : L₁.unop.finGaloisGroup ⟶ L₂.unop.finGaloisGroup := - haveI : Normal k L₂.unop := IsGalois.to_normal - letI : Algebra L₂.unop L₁.unop := RingHom.toAlgebra (Subsemiring.inclusion <| leOfHom le.1) - haveI : IsScalarTower k L₂.unop L₁.unop := IsScalarTower.of_algebraMap_eq (congrFun rfl) - FiniteGrp.ofHom (AlgEquiv.restrictNormalHom L₂.unop) - -namespace finGaloisGroupMap - -@[simp] -lemma map_id (L : (FiniteGaloisIntermediateField k K)ᵒᵖ) : - (finGaloisGroupMap (𝟙 L)) = 𝟙 L.unop.finGaloisGroup := - AlgEquiv.restrictNormalHom_id _ _ - -@[simp] -lemma map_comp {L₁ L₂ L₃ : (FiniteGaloisIntermediateField k K)ᵒᵖ} (f : L₁ ⟶ L₂) (g : L₂ ⟶ L₃) : - finGaloisGroupMap (f ≫ g) = finGaloisGroupMap f ≫ finGaloisGroupMap g := by - iterate 2 - induction L₁ with | _ L₁ => ?_ - induction L₂ with | _ L₂ => ?_ - induction L₃ with | _ L₃ => ?_ - letI : Algebra L₃ L₂ := RingHom.toAlgebra (Subsemiring.inclusion g.unop.le) - letI : Algebra L₂ L₁ := RingHom.toAlgebra (Subsemiring.inclusion f.unop.le) - letI : Algebra L₃ L₁ := RingHom.toAlgebra (Subsemiring.inclusion (g.unop.le.trans f.unop.le)) - haveI : IsScalarTower k L₂ L₁ := IsScalarTower.of_algebraMap_eq' rfl - haveI : IsScalarTower k L₃ L₁ := IsScalarTower.of_algebraMap_eq' rfl - haveI : IsScalarTower k L₃ L₂ := IsScalarTower.of_algebraMap_eq' rfl - haveI : IsScalarTower L₃ L₂ L₁ := IsScalarTower.of_algebraMap_eq' rfl - apply IsScalarTower.AlgEquiv.restrictNormalHom_comp k L₃ L₂ L₁ - -end finGaloisGroupMap - -variable (k K) in -/-- The functor from `FiniteGaloisIntermediateField` (ordered by reverse inclusion) to `FiniteGrp`, -mapping each intermediate field `K/L/k` to `Gal (L/k)`.-/ -noncomputable def finGaloisGroupFunctor : (FiniteGaloisIntermediateField k K)ᵒᵖ ⥤ FiniteGrp where - obj L := L.unop.finGaloisGroup - map := finGaloisGroupMap - map_id := finGaloisGroupMap.map_id - map_comp := finGaloisGroupMap.map_comp - -end Profinite diff --git a/Mathlib/FieldTheory/Galois/Profinite.lean b/Mathlib/FieldTheory/Galois/Profinite.lean new file mode 100644 index 0000000000000..300f56169316b --- /dev/null +++ b/Mathlib/FieldTheory/Galois/Profinite.lean @@ -0,0 +1,84 @@ +/- +Copyright (c) 2024 Nailin Guan. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Nailin Guan, Yuyang Zhao, Jujian Zhang +-/ + +import Mathlib.Algebra.Category.Grp.FiniteGrp +import Mathlib.CategoryTheory.Category.Preorder +import Mathlib.FieldTheory.Galois.GaloisClosure + +/-! + +# Main definitions and results + +In a field extension `K/k` + +* `finGaloisGroup L` : The (finite) Galois group `Gal(L/k)` associated to a + `L : FiniteGaloisIntermediateField k K` `L`. + +* `finGaloisGroupMap` : For `FiniteGaloisIntermediateField` s `L₁` and `L₂` with `L₂ ≤ L₁` + giving the restriction of `Gal(L₁/k)` to `Gal(L₂/k)` + +* `finGaloisGroupFunctor` : Mapping `FiniteGaloisIntermediateField` ordered by inverse inclusion to + its corresponding Galois Group as `FiniteGrp`. + +-/ + +open CategoryTheory Opposite + +variable {k K : Type*} [Field k] [Field K] [Algebra k K] + +section Profinite + +/-- The (finite) Galois group `Gal(L / k)` associated to a +`L : FiniteGaloisIntermediateField k K` `L`. -/ +def FiniteGaloisIntermediateField.finGaloisGroup (L : FiniteGaloisIntermediateField k K) : + FiniteGrp := + letI := AlgEquiv.fintype k L + FiniteGrp.of <| L ≃ₐ[k] L + +/-- For `FiniteGaloisIntermediateField` s `L₁` and `L₂` with `L₂ ≤ L₁` + the restriction homomorphism from `Gal(L₁/k)` to `Gal(L₂/k)` -/ +noncomputable def finGaloisGroupMap {L₁ L₂ : (FiniteGaloisIntermediateField k K)ᵒᵖ} + (le : L₁ ⟶ L₂) : L₁.unop.finGaloisGroup ⟶ L₂.unop.finGaloisGroup := + haveI : Normal k L₂.unop := IsGalois.to_normal + letI : Algebra L₂.unop L₁.unop := RingHom.toAlgebra (Subsemiring.inclusion <| leOfHom le.1) + haveI : IsScalarTower k L₂.unop L₁.unop := IsScalarTower.of_algebraMap_eq (congrFun rfl) + FiniteGrp.ofHom (AlgEquiv.restrictNormalHom L₂.unop) + +namespace finGaloisGroupMap + +@[simp] +lemma map_id (L : (FiniteGaloisIntermediateField k K)ᵒᵖ) : + (finGaloisGroupMap (𝟙 L)) = 𝟙 L.unop.finGaloisGroup := + AlgEquiv.restrictNormalHom_id _ _ + +@[simp] +lemma map_comp {L₁ L₂ L₃ : (FiniteGaloisIntermediateField k K)ᵒᵖ} (f : L₁ ⟶ L₂) (g : L₂ ⟶ L₃) : + finGaloisGroupMap (f ≫ g) = finGaloisGroupMap f ≫ finGaloisGroupMap g := by + iterate 2 + induction L₁ with | _ L₁ => ?_ + induction L₂ with | _ L₂ => ?_ + induction L₃ with | _ L₃ => ?_ + letI : Algebra L₃ L₂ := RingHom.toAlgebra (Subsemiring.inclusion g.unop.le) + letI : Algebra L₂ L₁ := RingHom.toAlgebra (Subsemiring.inclusion f.unop.le) + letI : Algebra L₃ L₁ := RingHom.toAlgebra (Subsemiring.inclusion (g.unop.le.trans f.unop.le)) + haveI : IsScalarTower k L₂ L₁ := IsScalarTower.of_algebraMap_eq' rfl + haveI : IsScalarTower k L₃ L₁ := IsScalarTower.of_algebraMap_eq' rfl + haveI : IsScalarTower k L₃ L₂ := IsScalarTower.of_algebraMap_eq' rfl + haveI : IsScalarTower L₃ L₂ L₁ := IsScalarTower.of_algebraMap_eq' rfl + apply IsScalarTower.AlgEquiv.restrictNormalHom_comp k L₃ L₂ L₁ + +end finGaloisGroupMap + +variable (k K) in +/-- The functor from `FiniteGaloisIntermediateField` (ordered by reverse inclusion) to `FiniteGrp`, +mapping each intermediate field `K/L/k` to `Gal (L/k)`.-/ +noncomputable def finGaloisGroupFunctor : (FiniteGaloisIntermediateField k K)ᵒᵖ ⥤ FiniteGrp where + obj L := L.unop.finGaloisGroup + map := finGaloisGroupMap + map_id := finGaloisGroupMap.map_id + map_comp := finGaloisGroupMap.map_comp + +end Profinite From 7ca6dfcd941c73f5993541abab8addd016f5208c Mon Sep 17 00:00:00 2001 From: leanprover-community-bot-assistant Date: Sun, 20 Oct 2024 00:15:00 +0000 Subject: [PATCH 195/505] chore(scripts): update nolints.json (#17958) I am happy to remove some nolints for you! --- scripts/nolints.json | 1 - 1 file changed, 1 deletion(-) diff --git a/scripts/nolints.json b/scripts/nolints.json index 8b556db384b5e..438084797d9f2 100644 --- a/scripts/nolints.json +++ b/scripts/nolints.json @@ -502,7 +502,6 @@ ["docBlame", "Mathlib.Tactic.tacticMatch_target_"], ["docBlame", "Mathlib.Tactic.tacticSet!_"], ["docBlame", "Mathlib.Tactic.tacticSuffices_"], - ["docBlame", "Mathlib.Tactic.tacticTransitivity___"], ["docBlame", "Mathlib.Tactic.usingArg"], ["docBlame", "Mathlib.Tactic.withArgs"], ["docBlame", "Mathlib.WhatsNew.diffExtension"], From 17ca754012e0011ae092c49c91174c1cbe11a007 Mon Sep 17 00:00:00 2001 From: FR Date: Sun, 20 Oct 2024 07:05:32 +0000 Subject: [PATCH 196/505] =?UTF-8?q?chore:=20avoid=20some=20`(0=20<=20((?= =?UTF-8?q?=E2=86=91)=20:=20=E2=84=9D=E2=89=A50=20=E2=86=92=20=E2=84=9D)?= =?UTF-8?q?=20x)=20=3D=20(0=20<=20x)`=20defeq=20abuses=20(#17946)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean | 6 +++--- Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean | 4 ++-- Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean | 2 +- 3 files changed, 6 insertions(+), 6 deletions(-) diff --git a/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean b/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean index 8164bbd406fee..31f90aefe8f74 100644 --- a/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean +++ b/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean @@ -535,9 +535,9 @@ theorem eLpNorm_le_eLpNorm_fderiv_of_eq_inner {u : E → F'} by_cases h3u : ∫⁻ x, ‖u x‖₊ ^ (p' : ℝ) ∂μ = 0 · rw [eLpNorm_nnreal_eq_lintegral h0p', h3u, ENNReal.zero_rpow_of_pos] <;> positivity have h4u : ∫⁻ x, ‖u x‖₊ ^ (p' : ℝ) ∂μ ≠ ∞ := by - refine lintegral_rpow_nnnorm_lt_top_of_eLpNorm'_lt_top (pos_iff_ne_zero.mpr h0p') ?_ |>.ne - dsimp only - rw [NNReal.val_eq_coe, ← eLpNorm_nnreal_eq_eLpNorm' h0p'] + refine lintegral_rpow_nnnorm_lt_top_of_eLpNorm'_lt_top + ((NNReal.coe_pos.trans pos_iff_ne_zero).mpr h0p') ?_ |>.ne + rw [← eLpNorm_nnreal_eq_eLpNorm' h0p'] exact hu.continuous.memℒp_of_hasCompactSupport (μ := μ) h2u |>.eLpNorm_lt_top have h5u : (∫⁻ x, ‖u x‖₊ ^ (p' : ℝ) ∂μ) ^ (1 / q) ≠ 0 := ENNReal.rpow_pos (pos_iff_ne_zero.mpr h3u) h4u |>.ne' diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean b/Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean index d2306464d57ff..acdcc41b5be97 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean @@ -79,7 +79,7 @@ theorem one_rpow (x : ℝ) : (1 : ℝ≥0) ^ x = 1 := NNReal.eq <| Real.one_rpow _ theorem rpow_add {x : ℝ≥0} (hx : x ≠ 0) (y z : ℝ) : x ^ (y + z) = x ^ y * x ^ z := - NNReal.eq <| Real.rpow_add (pos_iff_ne_zero.2 hx) _ _ + NNReal.eq <| Real.rpow_add ((NNReal.coe_pos.trans pos_iff_ne_zero).mpr hx) _ _ theorem rpow_add' (h : y + z ≠ 0) (x : ℝ≥0) : x ^ (y + z) = x ^ y * x ^ z := NNReal.eq <| Real.rpow_add' x.2 h @@ -146,7 +146,7 @@ lemma rpow_mul_intCast (x : ℝ≥0) (y : ℝ) (n : ℤ) : x ^ (y * n) = (x ^ y) theorem rpow_neg_one (x : ℝ≥0) : x ^ (-1 : ℝ) = x⁻¹ := by simp [rpow_neg] theorem rpow_sub {x : ℝ≥0} (hx : x ≠ 0) (y z : ℝ) : x ^ (y - z) = x ^ y / x ^ z := - NNReal.eq <| Real.rpow_sub (pos_iff_ne_zero.2 hx) y z + NNReal.eq <| Real.rpow_sub ((NNReal.coe_pos.trans pos_iff_ne_zero).mpr hx) y z theorem rpow_sub' (h : y - z ≠ 0) (x : ℝ≥0) : x ^ (y - z) = x ^ y / x ^ z := NNReal.eq <| Real.rpow_sub' x.2 h diff --git a/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean b/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean index 5bda6afb8f9e1..ee75398764b6e 100644 --- a/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean +++ b/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean @@ -139,7 +139,7 @@ alias lintegral_rpow_nnnorm_eq_rpow_snorm' := lintegral_rpow_nnnorm_eq_rpow_eLpN lemma eLpNorm_nnreal_pow_eq_lintegral {f : α → F} {p : ℝ≥0} (hp : p ≠ 0) : eLpNorm f p μ ^ (p : ℝ) = ∫⁻ x, ‖f x‖₊ ^ (p : ℝ) ∂μ := by simp [eLpNorm_eq_eLpNorm' (by exact_mod_cast hp) ENNReal.coe_ne_top, - lintegral_rpow_nnnorm_eq_rpow_eLpNorm' (show 0 < (p : ℝ) from pos_iff_ne_zero.mpr hp)] + lintegral_rpow_nnnorm_eq_rpow_eLpNorm' ((NNReal.coe_pos.trans pos_iff_ne_zero).mpr hp)] @[deprecated (since := "2024-07-27")] alias snorm_nnreal_pow_eq_lintegral := eLpNorm_nnreal_pow_eq_lintegral From 2f60fa67d7c9457a0e591806fc3201269c401030 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 20 Oct 2024 07:37:59 +0000 Subject: [PATCH 197/505] =?UTF-8?q?feat(Kernel):=20`map=20(=CE=BA=20=C3=97?= =?UTF-8?q?=E2=82=96=20=CE=B7)=20Prod.swap=20=3D=20=CE=B7=20=C3=97?= =?UTF-8?q?=E2=82=96=20=CE=BA`=20(#17920)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit and similar lemmas. From PFR Co-authored-by: Rémy Degenne --- Mathlib/Probability/Kernel/Composition.lean | 92 ++++++++++++++----- .../Kernel/Disintegration/Basic.lean | 2 +- .../Kernel/Disintegration/CDFToKernel.lean | 2 +- .../Kernel/Disintegration/Integral.lean | 4 +- .../Kernel/Disintegration/StandardBorel.lean | 2 +- .../Kernel/Disintegration/Unique.lean | 2 +- .../Probability/Kernel/IntegralCompProd.lean | 2 +- .../Probability/Kernel/MeasureCompProd.lean | 2 +- 8 files changed, 79 insertions(+), 29 deletions(-) diff --git a/Mathlib/Probability/Kernel/Composition.lean b/Mathlib/Probability/Kernel/Composition.lean index fd42f53736b76..9a2db5e2be4bc 100644 --- a/Mathlib/Probability/Kernel/Composition.lean +++ b/Mathlib/Probability/Kernel/Composition.lean @@ -208,8 +208,8 @@ theorem compProd_of_not_isSFiniteKernel_right (κ : Kernel α β) (η : Kernel ( rw [compProd, dif_neg] simp [h] -theorem compProd_apply (κ : Kernel α β) [IsSFiniteKernel κ] (η : Kernel (α × β) γ) - [IsSFiniteKernel η] (a : α) (hs : MeasurableSet s) : +theorem compProd_apply (hs : MeasurableSet s) (κ : Kernel α β) [IsSFiniteKernel κ] + (η : Kernel (α × β) γ) [IsSFiniteKernel η] (a : α) : (κ ⊗ₖ η) a s = ∫⁻ b, η (a, b) {c | (b, c) ∈ s} ∂κ a := compProd_apply_eq_compProdFun κ η a hs @@ -229,7 +229,7 @@ lemma compProd_zero_left (κ : Kernel (α × β) γ) : (0 : Kernel α β) ⊗ₖ κ = 0 := by by_cases h : IsSFiniteKernel κ · ext a s hs - rw [Kernel.compProd_apply _ _ _ hs] + rw [Kernel.compProd_apply hs] simp · rw [Kernel.compProd_of_not_isSFiniteKernel_right _ _ h] @@ -238,10 +238,42 @@ lemma compProd_zero_right (κ : Kernel α β) (γ : Type*) {mγ : MeasurableSpac κ ⊗ₖ (0 : Kernel (α × β) γ) = 0 := by by_cases h : IsSFiniteKernel κ · ext a s hs - rw [Kernel.compProd_apply _ _ _ hs] + rw [Kernel.compProd_apply hs] simp · rw [Kernel.compProd_of_not_isSFiniteKernel_left _ _ h] +lemma compProd_preimage_fst {s : Set β} (hs : MeasurableSet s) (κ : Kernel α β) + (η : Kernel (α × β) γ) [IsSFiniteKernel κ] [IsMarkovKernel η] (x : α) : + (κ ⊗ₖ η) x (Prod.fst ⁻¹' s) = κ x s := by + rw [compProd_apply (measurable_fst hs)] + simp only [Set.mem_preimage] + classical + have : ∀ b : β, η (x, b) {_c | b ∈ s} = s.indicator (fun _ ↦ 1) b := by + intro b + by_cases hb : b ∈ s <;> simp [hb] + simp_rw [this] + rw [lintegral_indicator_const hs, one_mul] + +lemma compProd_deterministic_apply [MeasurableSingletonClass γ] {f : α × β → γ} (hf : Measurable f) + {s : Set (β × γ)} (hs : MeasurableSet s) (κ : Kernel α β) [IsSFiniteKernel κ] (x : α) : + (κ ⊗ₖ deterministic f hf) x s = κ x {b | (b, f (x, b)) ∈ s} := by + simp only [deterministic_apply, measurableSet_setOf, Set.mem_setOf_eq, Measure.dirac_apply, + Set.mem_setOf_eq, Set.indicator_apply, Pi.one_apply, compProd_apply hs] + let t := {b | (b, f (x, b)) ∈ s} + have ht : MeasurableSet t := (measurable_id.prod_mk (hf.comp measurable_prod_mk_left)) hs + rw [← lintegral_add_compl _ ht] + convert add_zero _ + · suffices ∀ b ∈ tᶜ, (if (b, f (x, b)) ∈ s then (1 : ℝ≥0∞) else 0) = 0 by + rw [setLIntegral_congr_fun ht.compl (ae_of_all _ this), lintegral_zero] + intro b hb + simp only [t, Set.mem_compl_iff, Set.mem_setOf_eq] at hb + simp [hb] + · suffices ∀ b ∈ t, (if (b, f (x, b)) ∈ s then (1 : ℝ≥0∞) else 0) = 1 by + rw [setLIntegral_congr_fun ht (ae_of_all _ this), setLIntegral_one] + intro b hb + simp only [t, Set.mem_setOf_eq] at hb + simp [hb] + section Ae /-! ### `ae` filter of the composition-product -/ @@ -257,14 +289,14 @@ theorem ae_kernel_lt_top (a : α) (h2s : (κ ⊗ₖ η) a s ≠ ∞) : have ht : MeasurableSet t := measurableSet_toMeasurable _ _ have h2t : (κ ⊗ₖ η) a t ≠ ∞ := by rwa [measure_toMeasurable] have ht_lt_top : ∀ᵐ b ∂κ a, η (a, b) (Prod.mk b ⁻¹' t) < ∞ := by - rw [Kernel.compProd_apply _ _ _ ht] at h2t + rw [Kernel.compProd_apply ht] at h2t exact ae_lt_top (Kernel.measurable_kernel_prod_mk_left' ht a) h2t filter_upwards [ht_lt_top] with b hb exact (this b).trans_lt hb theorem compProd_null (a : α) (hs : MeasurableSet s) : (κ ⊗ₖ η) a s = 0 ↔ (fun b => η (a, b) (Prod.mk b ⁻¹' s)) =ᵐ[κ a] 0 := by - rw [Kernel.compProd_apply _ _ _ hs, lintegral_eq_zero_iff] + rw [Kernel.compProd_apply hs, lintegral_eq_zero_iff] · rfl · exact Kernel.measurable_kernel_prod_mk_left' hs a @@ -308,8 +340,8 @@ variable {κ : Kernel α β} [IsSFiniteKernel κ] {η : Kernel (α × β) γ} [I theorem compProd_restrict {s : Set β} {t : Set γ} (hs : MeasurableSet s) (ht : MeasurableSet t) : Kernel.restrict κ hs ⊗ₖ Kernel.restrict η ht = Kernel.restrict (κ ⊗ₖ η) (hs.prod ht) := by ext a u hu - rw [compProd_apply _ _ _ hu, restrict_apply' _ _ _ hu, - compProd_apply _ _ _ (hu.inter (hs.prod ht))] + rw [compProd_apply hu, restrict_apply' _ _ _ hu, + compProd_apply (hu.inter (hs.prod ht))] simp only [Kernel.restrict_apply, Measure.restrict_apply' ht, Set.mem_inter_iff, Set.prod_mk_mem_set_prod_eq] have : @@ -383,7 +415,7 @@ theorem lintegral_compProd' (κ : Kernel α β) [IsSFiniteKernel κ] (η : Kerne simp (config := { unfoldPartialApp := true }) only [SimpleFunc.const_zero, SimpleFunc.coe_piecewise, SimpleFunc.coe_const, SimpleFunc.coe_zero, Set.piecewise_eq_indicator, Function.const, lintegral_indicator_const hs] - rw [compProd_apply κ η _ hs, ← lintegral_const_mul c _] + rw [compProd_apply hs, ← lintegral_const_mul c _] swap · exact (measurable_kernel_prod_mk_left ((measurable_fst.snd.prod_mk measurable_snd) hs)).comp measurable_prod_mk_left @@ -485,11 +517,8 @@ theorem compProd_eq_sum_compProd_right (κ : Kernel α β) (η : Kernel (α × rw [Kernel.sum_comm] instance IsMarkovKernel.compProd (κ : Kernel α β) [IsMarkovKernel κ] (η : Kernel (α × β) γ) - [IsMarkovKernel η] : IsMarkovKernel (κ ⊗ₖ η) := - ⟨fun a => - ⟨by - rw [compProd_apply κ η a MeasurableSet.univ] - simp only [Set.mem_univ, Set.setOf_true, measure_univ, lintegral_one]⟩⟩ + [IsMarkovKernel η] : IsMarkovKernel (κ ⊗ₖ η) where + isProbabilityMeasure a := ⟨by simp [compProd_apply]⟩ theorem compProd_apply_univ_le (κ : Kernel α β) (η : Kernel (α × β) γ) [IsFiniteKernel η] (a : α) : (κ ⊗ₖ η) a Set.univ ≤ κ a Set.univ * IsFiniteKernel.bound η := by @@ -497,7 +526,7 @@ theorem compProd_apply_univ_le (κ : Kernel α β) (η : Kernel (α × β) γ) [ swap · rw [compProd_of_not_isSFiniteKernel_left _ _ hκ] simp - rw [compProd_apply κ η a MeasurableSet.univ] + rw [compProd_apply .univ] simp only [Set.mem_univ, Set.setOf_true] let Cη := IsFiniteKernel.bound η calc @@ -530,13 +559,13 @@ instance IsSFiniteKernel.compProd (κ : Kernel α β) (η : Kernel (α × β) γ lemma compProd_add_left (μ κ : Kernel α β) (η : Kernel (α × β) γ) [IsSFiniteKernel μ] [IsSFiniteKernel κ] [IsSFiniteKernel η] : - (μ + κ) ⊗ₖ η = μ ⊗ₖ η + κ ⊗ₖ η := by ext _ _ hs; simp [compProd_apply _ _ _ hs] + (μ + κ) ⊗ₖ η = μ ⊗ₖ η + κ ⊗ₖ η := by ext _ _ hs; simp [compProd_apply hs] lemma compProd_add_right (μ : Kernel α β) (κ η : Kernel (α × β) γ) [IsSFiniteKernel μ] [IsSFiniteKernel κ] [IsSFiniteKernel η] : μ ⊗ₖ (κ + η) = μ ⊗ₖ κ + μ ⊗ₖ η := by ext a s hs - simp only [compProd_apply _ _ _ hs, coe_add, Pi.add_apply, Measure.coe_add] + simp only [compProd_apply hs, coe_add, Pi.add_apply, Measure.coe_add] rw [lintegral_add_left] exact measurable_kernel_prod_mk_left' hs a @@ -545,7 +574,7 @@ lemma comapRight_compProd_id_prod {δ : Type*} {mδ : MeasurableSpace δ} {f : δ → γ} (hf : MeasurableEmbedding f) : comapRight (κ ⊗ₖ η) (MeasurableEmbedding.id.prod_mk hf) = κ ⊗ₖ (comapRight η hf) := by ext a t ht - rw [comapRight_apply' _ _ _ ht, compProd_apply, compProd_apply _ _ _ ht] + rw [comapRight_apply' _ _ _ ht, compProd_apply, compProd_apply ht] swap; · exact (MeasurableEmbedding.id.prod_mk hf).measurableSet_image.mpr ht refine lintegral_congr (fun b ↦ ?_) simp only [id_eq, Set.mem_image, Prod.mk.injEq, Prod.exists] @@ -1136,8 +1165,7 @@ section Prod /-! ### Product of two kernels -/ - -variable {γ : Type*} {mγ : MeasurableSpace γ} +variable {γ δ : Type*} {mγ : MeasurableSpace γ} {mδ : MeasurableSpace δ} /-- Product of two kernels. This is meaningful only when the kernels are s-finite. -/ noncomputable def prod (κ : Kernel α β) (η : Kernel α γ) : Kernel α (β × γ) := @@ -1148,7 +1176,7 @@ scoped[ProbabilityTheory] infixl:100 " ×ₖ " => ProbabilityTheory.Kernel.prod theorem prod_apply' (κ : Kernel α β) [IsSFiniteKernel κ] (η : Kernel α γ) [IsSFiniteKernel η] (a : α) {s : Set (β × γ)} (hs : MeasurableSet s) : (κ ×ₖ η) a s = ∫⁻ b : β, (η a) {c : γ | (b, c) ∈ s} ∂κ a := by - simp_rw [prod, compProd_apply _ _ _ hs, swapLeft_apply _ _, prodMkLeft_apply, Prod.swap_prod_mk] + simp_rw [prod, compProd_apply hs, swapLeft_apply _ _, prodMkLeft_apply, Prod.swap_prod_mk] lemma prod_apply (κ : Kernel α β) [IsSFiniteKernel κ] (η : Kernel α γ) [IsSFiniteKernel η] (a : α) : @@ -1192,6 +1220,28 @@ instance IsSFiniteKernel.prod (κ : Kernel α β) (η : Kernel α γ) : snd (κ ×ₖ η) = η := by ext x; simp [snd_apply, prod_apply] +lemma comap_prod_swap (κ : Kernel α β) (η : Kernel γ δ) [IsSFiniteKernel κ] [IsSFiniteKernel η] : + comap (prodMkRight α η ×ₖ prodMkLeft γ κ) Prod.swap measurable_swap + = map (prodMkRight γ κ ×ₖ prodMkLeft α η) Prod.swap := by + rw [ext_fun_iff] + intro x f hf + rw [lintegral_comap, lintegral_map _ measurable_swap _ hf, lintegral_prod _ _ _ hf, + lintegral_prod] + swap; · exact hf.comp measurable_swap + simp only [prodMkRight_apply, Prod.fst_swap, Prod.swap_prod_mk, lintegral_prodMkLeft, + Prod.snd_swap] + refine (lintegral_lintegral_swap ?_).symm + exact (hf.comp measurable_swap).aemeasurable + +lemma map_prod_swap (κ : Kernel α β) (η : Kernel α γ) [IsSFiniteKernel κ] [IsSFiniteKernel η] : + map (κ ×ₖ η) Prod.swap = η ×ₖ κ := by + rw [ext_fun_iff] + intro x f hf + rw [lintegral_map _ measurable_swap _ hf, lintegral_prod, lintegral_prod _ _ _ hf] + swap; · exact hf.comp measurable_swap + refine (lintegral_lintegral_swap ?_).symm + exact hf.aemeasurable + lemma deterministic_prod_deterministic {f : α → β} {g : α → γ} (hf : Measurable f) (hg : Measurable g) : deterministic f hf ×ₖ deterministic g hg diff --git a/Mathlib/Probability/Kernel/Disintegration/Basic.lean b/Mathlib/Probability/Kernel/Disintegration/Basic.lean index 85c938f894937..22934beefaf08 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Basic.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Basic.lean @@ -163,7 +163,7 @@ instance condKernelCountable.instIsCondKernel [∀ a, IsMarkovKernel (κCond a)] constructor ext a s hs conv_rhs => rw [← (κ a).disintegrate (κCond a)] - simp_rw [compProd_apply _ _ _ hs, condKernelCountable_apply, Measure.compProd_apply hs] + simp_rw [compProd_apply hs, condKernelCountable_apply, Measure.compProd_apply hs] congr end Countable diff --git a/Mathlib/Probability/Kernel/Disintegration/CDFToKernel.lean b/Mathlib/Probability/Kernel/Disintegration/CDFToKernel.lean index 3655fb84066ea..cf2267887837b 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CDFToKernel.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CDFToKernel.lean @@ -684,7 +684,7 @@ lemma lintegral_toKernel_mem [IsFiniteKernel κ] (hf : IsCondKernelCDF f κ ν) lemma compProd_toKernel [IsFiniteKernel κ] [IsSFiniteKernel ν] (hf : IsCondKernelCDF f κ ν) : ν ⊗ₖ hf.toKernel f = κ := by ext a s hs - rw [Kernel.compProd_apply _ _ _ hs, lintegral_toKernel_mem hf a hs] + rw [Kernel.compProd_apply hs, lintegral_toKernel_mem hf a hs] end diff --git a/Mathlib/Probability/Kernel/Disintegration/Integral.lean b/Mathlib/Probability/Kernel/Disintegration/Integral.lean index 6f9b34c32006f..b3b312597fb26 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Integral.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Integral.lean @@ -39,14 +39,14 @@ variable [CountableOrCountablyGenerated α β] {κ : Kernel α (β × Ω)} [IsFi lemma lintegral_condKernel_mem (a : α) {s : Set (β × Ω)} (hs : MeasurableSet s) : ∫⁻ x, Kernel.condKernel κ (a, x) {y | (x, y) ∈ s} ∂(Kernel.fst κ a) = κ a s := by conv_rhs => rw [← κ.disintegrate κ.condKernel] - simp_rw [Kernel.compProd_apply _ _ _ hs] + simp_rw [Kernel.compProd_apply hs] lemma setLIntegral_condKernel_eq_measure_prod (a : α) {s : Set β} (hs : MeasurableSet s) {t : Set Ω} (ht : MeasurableSet t) : ∫⁻ b in s, Kernel.condKernel κ (a, b) t ∂(Kernel.fst κ a) = κ a (s ×ˢ t) := by have : κ a (s ×ˢ t) = (Kernel.fst κ ⊗ₖ Kernel.condKernel κ) a (s ×ˢ t) := by congr; exact (κ.disintegrate _).symm - rw [this, Kernel.compProd_apply _ _ _ (hs.prod ht)] + rw [this, Kernel.compProd_apply (hs.prod ht)] classical have : ∀ b, Kernel.condKernel κ (a, b) {c | (b, c) ∈ s ×ˢ t} = s.indicator (fun b ↦ Kernel.condKernel κ (a, b) t) b := by diff --git a/Mathlib/Probability/Kernel/Disintegration/StandardBorel.lean b/Mathlib/Probability/Kernel/Disintegration/StandardBorel.lean index db4a4a141b457..b6934d43877ad 100644 --- a/Mathlib/Probability/Kernel/Disintegration/StandardBorel.lean +++ b/Mathlib/Probability/Kernel/Disintegration/StandardBorel.lean @@ -259,7 +259,7 @@ lemma compProd_fst_borelMarkovFromReal_eq_comapRight_compProd congr rw [h_fst] ext a t ht : 2 - simp_rw [compProd_apply _ _ _ ht] + simp_rw [compProd_apply ht] refine lintegral_congr_ae ?_ have h_ae : ∀ᵐ t ∂(fst κ a), (a, t) ∈ {p : α × β | η p (range e)ᶜ = 0} := by rw [← h_fst] diff --git a/Mathlib/Probability/Kernel/Disintegration/Unique.lean b/Mathlib/Probability/Kernel/Disintegration/Unique.lean index 3a2d73da4ca97..786669c638b32 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Unique.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Unique.lean @@ -134,7 +134,7 @@ lemma Kernel.apply_eq_measure_condKernel_of_compProd_eq have : ρ a = (ρ a).fst ⊗ₘ Kernel.comap κ (fun b ↦ (a, b)) measurable_prod_mk_left := by ext s hs conv_lhs => rw [← hκ] - rw [Measure.compProd_apply hs, Kernel.compProd_apply _ _ _ hs] + rw [Measure.compProd_apply hs, Kernel.compProd_apply hs] rfl have h := eq_condKernel_of_measure_eq_compProd _ this rw [Kernel.fst_apply] diff --git a/Mathlib/Probability/Kernel/IntegralCompProd.lean b/Mathlib/Probability/Kernel/IntegralCompProd.lean index a8ad7e64ece6c..e88e7a0f5f916 100644 --- a/Mathlib/Probability/Kernel/IntegralCompProd.lean +++ b/Mathlib/Probability/Kernel/IntegralCompProd.lean @@ -234,7 +234,7 @@ theorem integral_compProd : rotate_left · exact (Kernel.measurable_kernel_prod_mk_left' hs _).aemeasurable · exact ae_kernel_lt_top a h2s.ne - rw [Kernel.compProd_apply _ _ _ hs] + rw [Kernel.compProd_apply hs] rfl · intro f g _ i_f i_g hf hg simp_rw [integral_add' i_f i_g, Kernel.integral_integral_add' i_f i_g, hf, hg] diff --git a/Mathlib/Probability/Kernel/MeasureCompProd.lean b/Mathlib/Probability/Kernel/MeasureCompProd.lean index 9e4b40f0b1773..b2e9841c037ba 100644 --- a/Mathlib/Probability/Kernel/MeasureCompProd.lean +++ b/Mathlib/Probability/Kernel/MeasureCompProd.lean @@ -55,7 +55,7 @@ lemma compProd_of_not_isSFiniteKernel (μ : Measure α) (κ : Kernel α β) (h : lemma compProd_apply [SFinite μ] [IsSFiniteKernel κ] {s : Set (α × β)} (hs : MeasurableSet s) : (μ ⊗ₘ κ) s = ∫⁻ a, κ a (Prod.mk a ⁻¹' s) ∂μ := by - simp_rw [compProd, Kernel.compProd_apply _ _ _ hs, Kernel.const_apply, Kernel.prodMkLeft_apply'] + simp_rw [compProd, Kernel.compProd_apply hs, Kernel.const_apply, Kernel.prodMkLeft_apply'] rfl lemma compProd_apply_prod [SFinite μ] [IsSFiniteKernel κ] From a555db73ee24eb1fe6ced83622d96b66e54d1d36 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sun, 20 Oct 2024 10:03:18 +0000 Subject: [PATCH 198/505] chore(Filter): move more defs to `Defs` (#17949) --- Mathlib/Order/Filter/Defs.lean | 10 ++++++++++ Mathlib/Order/Filter/Lift.lean | 10 ---------- 2 files changed, 10 insertions(+), 10 deletions(-) diff --git a/Mathlib/Order/Filter/Defs.lean b/Mathlib/Order/Filter/Defs.lean index 06e6369f2f301..4b9a6cb4f9a6b 100644 --- a/Mathlib/Order/Filter/Defs.lean +++ b/Mathlib/Order/Filter/Defs.lean @@ -340,6 +340,16 @@ instance : Bind Filter := instance : Functor Filter where map := @Filter.map +/-- A variant on `bind` using a function `g` taking a set instead of a member of `α`. +This is essentially a push-forward along a function mapping each set to a filter. -/ +protected def lift (f : Filter α) (g : Set α → Filter β) := + ⨅ s ∈ f, g s + +/-- Specialize `lift` to functions `Set α → Set β`. This can be viewed as a generalization of `map`. +This is essentially a push-forward along a function mapping each set to a set. -/ +protected def lift' (f : Filter α) (h : Set α → Set β) := + f.lift (𝓟 ∘ h) + end Filter namespace Mathlib.Tactic diff --git a/Mathlib/Order/Filter/Lift.lean b/Mathlib/Order/Filter/Lift.lean index 95e8d4708a1be..f485af84c1df0 100644 --- a/Mathlib/Order/Filter/Lift.lean +++ b/Mathlib/Order/Filter/Lift.lean @@ -19,11 +19,6 @@ variable {α β γ : Type*} {ι : Sort*} section lift -/-- A variant on `bind` using a function `g` taking a set instead of a member of `α`. -This is essentially a push-forward along a function mapping each set to a filter. -/ -protected def lift (f : Filter α) (g : Set α → Filter β) := - ⨅ s ∈ f, g s - variable {f f₁ f₂ : Filter α} {g g₁ g₂ : Set α → Filter β} @[simp] @@ -200,11 +195,6 @@ end lift section Lift' -/-- Specialize `lift` to functions `Set α → Set β`. This can be viewed as a generalization of `map`. -This is essentially a push-forward along a function mapping each set to a set. -/ -protected def lift' (f : Filter α) (h : Set α → Set β) := - f.lift (𝓟 ∘ h) - variable {f f₁ f₂ : Filter α} {h h₁ h₂ : Set α → Set β} @[simp] From 1ae9a84cb90c97164a22a64da2604a7324d0c9f6 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Sun, 20 Oct 2024 11:19:44 +0000 Subject: [PATCH 199/505] feat(AlgebraicGeometry): affine morphisms are separated (#17962) --- .../Morphisms/ClosedImmersion.lean | 42 ++++++++++++++++++- .../Morphisms/Separated.lean | 23 ++++++++-- Mathlib/AlgebraicGeometry/Pullbacks.lean | 13 ++++++ 3 files changed, 72 insertions(+), 6 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean b/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean index 03d557fa0c469..0e01ff0984171 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean @@ -3,10 +3,12 @@ Copyright (c) 2023 Jonas van der Schaaf. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Amelia Livingston, Christian Merten, Jonas van der Schaaf -/ -import Mathlib.AlgebraicGeometry.Morphisms.QuasiSeparated +import Mathlib.AlgebraicGeometry.Morphisms.Affine import Mathlib.AlgebraicGeometry.Morphisms.RingHomProperties import Mathlib.AlgebraicGeometry.Morphisms.FiniteType +import Mathlib.AlgebraicGeometry.Morphisms.IsIso import Mathlib.AlgebraicGeometry.ResidueField +import Mathlib.AlgebraicGeometry.Properties /-! @@ -110,7 +112,7 @@ instance spec_of_quotient_mk {R : CommRingCat.{u}} (I : Ideal R) : /-- Any morphism between affine schemes that is surjective on global sections is a closed immersion. -/ lemma of_surjective_of_isAffine {X Y : Scheme} [IsAffine X] [IsAffine Y] (f : X ⟶ Y) - (h : Function.Surjective (Scheme.Γ.map f.op)) : IsClosedImmersion f := by + (h : Function.Surjective (f.app ⊤)) : IsClosedImmersion f := by rw [MorphismProperty.arrow_mk_iso_iff @IsClosedImmersion (arrowIsoSpecΓOfIsAffine f)] apply spec_of_surjective exact h @@ -262,6 +264,17 @@ instance IsClosedImmersion.hasAffineProperty : HasAffineProperty @IsClosedImmers convert HasAffineProperty.of_isLocalAtTarget @IsClosedImmersion refine ⟨fun ⟨h₁, h₂⟩ ↦ of_surjective_of_isAffine _ h₂, by apply isAffine_surjective_of_isAffine⟩ +instance (priority := 900) {X Y : Scheme.{u}} (f : X ⟶ Y) [h : IsClosedImmersion f] : + IsAffineHom f := by + wlog hY : IsAffine Y + · rw [IsLocalAtTarget.iff_of_iSup_eq_top (P := @IsAffineHom) _ + (iSup_affineOpens_eq_top Y)] + intro U + have H : IsClosedImmersion (f ∣_ U) := IsLocalAtTarget.restrict h U + exact this _ U.2 + rw [HasAffineProperty.iff_of_isAffine (P := @IsAffineHom)] + exact (IsClosedImmersion.isAffine_surjective_of_isAffine f).1 + /-- Being a closed immersion is stable under base change. -/ lemma IsClosedImmersion.stableUnderBaseChange : MorphismProperty.StableUnderBaseChange @IsClosedImmersion := by @@ -285,4 +298,29 @@ instance (priority := 900) {X Y : Scheme.{u}} (f : X ⟶ Y) [h : IsClosedImmersi rw [HasRingHomProperty.iff_of_isAffine (P := @LocallyOfFiniteType)] exact RingHom.FiniteType.of_surjective (Scheme.Hom.app f ⊤) hf +/-- A surjective closed immersion is an isomorphism when the target is reduced. -/ +lemma isIso_of_isClosedImmersion_of_surjective {X Y : Scheme.{u}} (f : X ⟶ Y) + [IsClosedImmersion f] [Surjective f] [IsReduced Y] : + IsIso f := by + wlog hY : IsAffine Y + · refine (IsLocalAtTarget.iff_of_openCover (P := .isomorphisms Scheme) Y.affineCover).mpr ?_ + intro i + apply (config := { allowSynthFailures := true }) this + · exact IsClosedImmersion.stableUnderBaseChange.snd _ _ inferInstance + · exact IsLocalAtTarget.of_isPullback (.of_hasPullback f (Y.affineCover.map i)) ‹_› + · exact isReduced_of_isOpenImmersion (Y.affineCover.map i) + · infer_instance + apply IsClosedImmersion.isIso_of_injective_of_isAffine + obtain ⟨hX, hf⟩ := HasAffineProperty.iff_of_isAffine.mp ‹IsClosedImmersion f› + let φ := f.app ⊤ + suffices RingHom.ker φ ≤ nilradical _ by + rwa [nilradical_eq_zero, Submodule.zero_eq_bot, le_bot_iff, + ← RingHom.injective_iff_ker_eq_bot] at this + refine (PrimeSpectrum.zeroLocus_eq_top_iff _).mp ?_ + rw [← range_specComap_of_surjective _ _ hf, Set.top_eq_univ, Set.range_iff_surjective] + have : Surjective (Spec.map (f.app ⊤)) := + (MorphismProperty.arrow_mk_iso_iff @Surjective (arrowIsoSpecΓOfIsAffine f)).mp + (inferInstanceAs (Surjective f)) + exact this.1 + end AlgebraicGeometry diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Separated.lean b/Mathlib/AlgebraicGeometry/Morphisms/Separated.lean index 369855b7c262d..a1d81cfa22536 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Separated.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Separated.lean @@ -14,10 +14,6 @@ import Mathlib.CategoryTheory.MorphismProperty.Limits A morphism of schemes is separated if its diagonal morphism is a closed immmersion. -## TODO - -- Show affine morphisms are separated. - -/ @@ -75,6 +71,25 @@ instance : IsLocalAtTarget @IsSeparated := by rw [isSeparated_eq_diagonal_isClosedImmersion] infer_instance +instance (R S : CommRingCat.{u}) (f : R ⟶ S) : IsSeparated (Spec.map f) := by + constructor + letI := f.toAlgebra + show IsClosedImmersion (Limits.pullback.diagonal (Spec.map (CommRingCat.ofHom (algebraMap R S)))) + rw [diagonal_Spec_map, MorphismProperty.cancel_right_of_respectsIso @IsClosedImmersion] + exact .spec_of_surjective _ fun x ↦ ⟨.tmul R 1 x, + (Algebra.TensorProduct.lmul'_apply_tmul (R := R) (S := S) 1 x).trans (one_mul x)⟩ + +instance (priority := 100) [h : IsAffineHom f] : IsSeparated f := by + wlog hY : IsAffine Y + · rw [IsLocalAtTarget.iff_of_iSup_eq_top (P := @IsSeparated) _ + (iSup_affineOpens_eq_top Y)] + intro U + have H : IsAffineHom (f ∣_ U) := IsLocalAtTarget.restrict h U + exact this _ U.2 + have : IsAffine X := HasAffineProperty.iff_of_isAffine.mp h + rw [MorphismProperty.arrow_mk_iso_iff @IsSeparated (arrowIsoSpecΓOfIsAffine f)] + infer_instance + end IsSeparated end AlgebraicGeometry diff --git a/Mathlib/AlgebraicGeometry/Pullbacks.lean b/Mathlib/AlgebraicGeometry/Pullbacks.lean index ca37134839d14..712933f44210e 100644 --- a/Mathlib/AlgebraicGeometry/Pullbacks.lean +++ b/Mathlib/AlgebraicGeometry/Pullbacks.lean @@ -565,6 +565,7 @@ def pullbackSpecIso : letI H := IsLimit.equivIsoLimit (PullbackCone.eta _) (PushoutCocone.isColimitEquivIsLimitOp _ (CommRingCat.pushoutCoconeIsColimit R S T)) limit.isoLimitCone ⟨_, isLimitPullbackConeMapOfIsLimit Scheme.Spec _ H⟩ + /-- The composition of the inverse of the isomorphism `pullbackSepcIso R S T` (from the pullback of `Spec S ⟶ Spec R` and `Spec T ⟶ Spec R` to `Spec (S ⊗[R] T)`) with the first projection is @@ -575,6 +576,7 @@ the morphism `Spec (S ⊗[R] T) ⟶ Spec S` obtained by applying `Spec.map` to t lemma pullbackSpecIso_inv_fst : (pullbackSpecIso R S T).inv ≫ pullback.fst _ _ = Spec.map (ofHom includeLeftRingHom) := limit.isoLimitCone_inv_π _ _ + /-- The composition of the inverse of the isomorphism `pullbackSepcIso R S T` (from the pullback of `Spec S ⟶ Spec R` and `Spec T ⟶ Spec R` to `Spec (S ⊗[R] T)`) with the second projection is @@ -586,6 +588,7 @@ lemma pullbackSpecIso_inv_snd : (pullbackSpecIso R S T).inv ≫ pullback.snd _ _ = Spec.map (ofHom (R := T) (S := S ⊗[R] T) (toRingHom includeRight)) := limit.isoLimitCone_inv_π _ _ + /-- The composition of the isomorphism `pullbackSepcIso R S T` (from the pullback of `Spec S ⟶ Spec R` and `Spec T ⟶ Spec R` to `Spec (S ⊗[R] T)`) with the morphism @@ -596,6 +599,7 @@ is the first projection. lemma pullbackSpecIso_hom_fst : (pullbackSpecIso R S T).hom ≫ Spec.map (ofHom includeLeftRingHom) = pullback.fst _ _ := by rw [← pullbackSpecIso_inv_fst, Iso.hom_inv_id_assoc] + /-- The composition of the isomorphism `pullbackSepcIso R S T` (from the pullback of `Spec S ⟶ Spec R` and `Spec T ⟶ Spec R` to `Spec (S ⊗[R] T)`) with the morphism @@ -607,6 +611,15 @@ lemma pullbackSpecIso_hom_snd : (pullbackSpecIso R S T).hom ≫ Spec.map (ofHom (toRingHom includeRight)) = pullback.snd _ _ := by rw [← pullbackSpecIso_inv_snd, Iso.hom_inv_id_assoc] +lemma diagonal_Spec_map : + pullback.diagonal (Spec.map (CommRingCat.ofHom (algebraMap R S))) = + Spec.map (CommRingCat.ofHom (Algebra.TensorProduct.lmul' R : S ⊗[R] S →ₐ[R] S).toRingHom) ≫ + (pullbackSpecIso R S S).inv := by + ext1 <;> simp only [pullback.diagonal_fst, pullback.diagonal_snd, ← Spec.map_comp, ← Spec.map_id, + AlgHom.toRingHom_eq_coe, Category.assoc, pullbackSpecIso_inv_fst, pullbackSpecIso_inv_snd] + · congr 1; ext x; show x = Algebra.TensorProduct.lmul' R (S := S) (x ⊗ₜ[R] 1); simp + · congr 1; ext x; show x = Algebra.TensorProduct.lmul' R (S := S) (1 ⊗ₜ[R] x); simp + end Spec From 6a88b5460a66aeed2fb64195c0925fd4d07334f1 Mon Sep 17 00:00:00 2001 From: FR Date: Sun, 20 Oct 2024 12:27:18 +0000 Subject: [PATCH 200/505] chore: move binary recursion on `Nat` into a new file (#15567) This is #13649 with fewer definition changes. --- Mathlib.lean | 1 + Mathlib/Data/Nat/BinaryRec.lean | 157 +++++++++++++++++++++++++++++++ Mathlib/Data/Nat/Bits.lean | 140 +-------------------------- Mathlib/Data/Nat/Bitwise.lean | 10 +- Mathlib/Data/Nat/EvenOddRec.lean | 3 +- Mathlib/Data/Nat/Fib/Basic.lean | 3 +- Mathlib/Data/Num/Basic.lean | 6 +- Mathlib/Data/Tree/Get.lean | 1 + scripts/noshake.json | 4 +- 9 files changed, 172 insertions(+), 153 deletions(-) create mode 100644 Mathlib/Data/Nat/BinaryRec.lean diff --git a/Mathlib.lean b/Mathlib.lean index 3ace00fab33be..0eb264c536463 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2434,6 +2434,7 @@ import Mathlib.Data.NNRat.Lemmas import Mathlib.Data.NNRat.Order import Mathlib.Data.NNReal.Basic import Mathlib.Data.NNReal.Star +import Mathlib.Data.Nat.BinaryRec import Mathlib.Data.Nat.BitIndices import Mathlib.Data.Nat.Bits import Mathlib.Data.Nat.Bitwise diff --git a/Mathlib/Data/Nat/BinaryRec.lean b/Mathlib/Data/Nat/BinaryRec.lean new file mode 100644 index 0000000000000..0edc733fa68b0 --- /dev/null +++ b/Mathlib/Data/Nat/BinaryRec.lean @@ -0,0 +1,157 @@ +/- +Copyright (c) 2017 Microsoft Corporation. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Mario Carneiro, Praneeth Kolichala, Yuyang Zhao +-/ + +/-! +# Binary recursion on `Nat` + +This file defines binary recursion on `Nat`. + +## Main results +* `Nat.binaryRec`: A recursion principle for `bit` representations of natural numbers. +* `Nat.binaryRec'`: The same as `binaryRec`, but the induction step can assume that if `n=0`, + the bit being appended is `true`. +* `Nat.binaryRecFromOne`: The same as `binaryRec`, but special casing both 0 and 1 as base cases. +-/ + +universe u + +namespace Nat + +/-- `bit b` appends the digit `b` to the binary representation of its natural number input. -/ +def bit (b : Bool) : Nat → Nat := cond b (2 * · + 1) (2 * ·) + +theorem shiftRight_one (n) : n >>> 1 = n / 2 := rfl + +@[simp] +theorem bit_decide_mod_two_eq_one_shiftRight_one (n : Nat) : bit (n % 2 = 1) (n >>> 1) = n := by + simp only [bit, shiftRight_one] + cases mod_two_eq_zero_or_one n with | _ h => simpa [h] using Nat.div_add_mod n 2 + +theorem bit_testBit_zero_shiftRight_one (n : Nat) : bit (n.testBit 0) (n >>> 1) = n := by + simp + +@[simp] +theorem bit_eq_zero_iff {n : Nat} {b : Bool} : bit b n = 0 ↔ n = 0 ∧ b = false := by + cases n <;> cases b <;> simp [bit, Nat.shiftLeft_succ, Nat.two_mul, ← Nat.add_assoc] + +/-- For a predicate `motive : Nat → Sort u`, if instances can be + constructed for natural numbers of the form `bit b n`, + they can be constructed for any given natural number. -/ +@[inline] +def bitCasesOn {motive : Nat → Sort u} (n) (h : ∀ b n, motive (bit b n)) : motive n := + -- `1 &&& n != 0` is faster than `n.testBit 0`. This may change when we have faster `testBit`. + let x := h (1 &&& n != 0) (n >>> 1) + -- `congrArg motive _ ▸ x` is defeq to `x` in non-dependent case + congrArg motive n.bit_testBit_zero_shiftRight_one ▸ x + +/-- A recursion principle for `bit` representations of natural numbers. + For a predicate `motive : Nat → Sort u`, if instances can be + constructed for natural numbers of the form `bit b n`, + they can be constructed for all natural numbers. -/ +@[elab_as_elim, specialize] +def binaryRec {motive : Nat → Sort u} (z : motive 0) (f : ∀ b n, motive n → motive (bit b n)) + (n : Nat) : motive n := + if n0 : n = 0 then congrArg motive n0 ▸ z + else + let x := f (1 &&& n != 0) (n >>> 1) (binaryRec z f (n >>> 1)) + congrArg motive n.bit_testBit_zero_shiftRight_one ▸ x +decreasing_by exact bitwise_rec_lemma n0 + +/-- The same as `binaryRec`, but the induction step can assume that if `n=0`, + the bit being appended is `true`-/ +@[elab_as_elim, specialize] +def binaryRec' {motive : Nat → Sort u} (z : motive 0) + (f : ∀ b n, (n = 0 → b = true) → motive n → motive (bit b n)) : + ∀ n, motive n := + binaryRec z fun b n ih => + if h : n = 0 → b = true then f b n h ih + else + have : bit b n = 0 := by + rw [bit_eq_zero_iff] + cases n <;> cases b <;> simp at h ⊢ + congrArg motive this ▸ z + +/-- The same as `binaryRec`, but special casing both 0 and 1 as base cases -/ +@[elab_as_elim, specialize] +def binaryRecFromOne {motive : Nat → Sort u} (z₀ : motive 0) (z₁ : motive 1) + (f : ∀ b n, n ≠ 0 → motive n → motive (bit b n)) : + ∀ n, motive n := + binaryRec' z₀ fun b n h ih => + if h' : n = 0 then + have : bit b n = bit true 0 := by + rw [h', h h'] + congrArg motive this ▸ z₁ + else f b n h' ih + +theorem bit_val (b n) : bit b n = 2 * n + b.toNat := by + cases b <;> rfl + +@[simp] +theorem bit_div_two (b n) : bit b n / 2 = n := by + rw [bit_val, Nat.add_comm, add_mul_div_left, div_eq_of_lt, Nat.zero_add] + · cases b <;> decide + · decide + +@[simp] +theorem bit_mod_two (b n) : bit b n % 2 = b.toNat := by + cases b <;> simp [bit_val, mul_add_mod] + +@[simp] +theorem bit_shiftRight_one (b n) : bit b n >>> 1 = n := + bit_div_two b n + +theorem testBit_bit_zero (b n) : (bit b n).testBit 0 = b := by + simp + +@[simp] +theorem bitCasesOn_bit {motive : Nat → Sort u} (h : ∀ b n, motive (bit b n)) (b : Bool) (n : Nat) : + bitCasesOn (bit b n) h = h b n := by + change congrArg motive (bit b n).bit_testBit_zero_shiftRight_one ▸ h _ _ = h b n + generalize congrArg motive (bit b n).bit_testBit_zero_shiftRight_one = e; revert e + rw [testBit_bit_zero, bit_shiftRight_one] + intros; rfl + +unseal binaryRec in +@[simp] +theorem binaryRec_zero {motive : Nat → Sort u} (z : motive 0) (f : ∀ b n, motive n → motive (bit b n)) : + binaryRec z f 0 = z := + rfl + +@[simp] +theorem binaryRec_one {motive : Nat → Sort u} (z : motive 0) (f : ∀ b n, motive n → motive (bit b n)) : + binaryRec (motive := motive) z f 1 = f true 0 z := by + rw [binaryRec] + simp only [add_one_ne_zero, ↓reduceDIte, Nat.reduceShiftRight, binaryRec_zero] + rfl + +/-- +The same as `binaryRec_eq`, +but that one unfortunately requires `f` to be the identity when appending `false` to `0`. +Here, we allow you to explicitly say that that case is not happening, +i.e. supplying `n = 0 → b = true`. -/ +theorem binaryRec_eq' {motive : Nat → Sort u} {z : motive 0} + {f : ∀ b n, motive n → motive (bit b n)} (b n) + (h : f false 0 z = z ∨ (n = 0 → b = true)) : + binaryRec z f (bit b n) = f b n (binaryRec z f n) := by + by_cases h' : bit b n = 0 + case pos => + obtain ⟨rfl, rfl⟩ := bit_eq_zero_iff.mp h' + simp only [Bool.false_eq_true, imp_false, not_true_eq_false, or_false] at h + unfold binaryRec + exact h.symm + case neg => + rw [binaryRec, dif_neg h'] + change congrArg motive (bit b n).bit_testBit_zero_shiftRight_one ▸ f _ _ _ = _ + generalize congrArg motive (bit b n).bit_testBit_zero_shiftRight_one = e; revert e + rw [testBit_bit_zero, bit_shiftRight_one] + intros; rfl + +theorem binaryRec_eq {motive : Nat → Sort u} {z : motive 0} {f : ∀ b n, motive n → motive (bit b n)} + (h : f false 0 z = z) (b n) : + binaryRec z f (bit b n) = f b n (binaryRec z f n) := + binaryRec_eq' b n (.inl h) + +end Nat diff --git a/Mathlib/Data/Nat/Bits.lean b/Mathlib/Data/Nat/Bits.lean index 6ea51b387201e..bdcc8d87a6a7c 100644 --- a/Mathlib/Data/Nat/Bits.lean +++ b/Mathlib/Data/Nat/Bits.lean @@ -5,8 +5,8 @@ Authors: Praneeth Kolichala -/ import Mathlib.Algebra.Group.Basic import Mathlib.Algebra.Group.Nat -import Mathlib.Algebra.Group.Units.Basic import Mathlib.Data.Nat.Defs +import Mathlib.Data.Nat.BinaryRec import Mathlib.Data.List.Defs import Mathlib.Tactic.Convert import Mathlib.Tactic.GeneralizeProofs @@ -114,35 +114,9 @@ lemma div2_val (n) : div2 n = n / 2 := by (Nat.add_left_cancel (Eq.trans ?_ (Nat.mod_add_div n 2).symm)) rw [mod_two_of_bodd, bodd_add_div2] -/-- `bit b` appends the digit `b` to the binary representation of its natural number input. -/ -def bit (b : Bool) : ℕ → ℕ := cond b (2 * · + 1) (2 * ·) - -lemma bit_val (b n) : bit b n = 2 * n + b.toNat := by - cases b <;> rfl - lemma bit_decomp (n : Nat) : bit (bodd n) (div2 n) = n := (bit_val _ _).trans <| (Nat.add_comm _ _).trans <| bodd_add_div2 _ -theorem shiftRight_one (n) : n >>> 1 = n / 2 := rfl - -@[simp] -theorem bit_decide_mod_two_eq_one_shiftRight_one (n : Nat) : bit (n % 2 = 1) (n >>> 1) = n := by - simp only [bit, shiftRight_one] - cases mod_two_eq_zero_or_one n with | _ h => simpa [h] using Nat.div_add_mod n 2 - -theorem bit_testBit_zero_shiftRight_one (n : Nat) : bit (n.testBit 0) (n >>> 1) = n := by - simp - -/-- For a predicate `motive : Nat → Sort*`, if instances can be - constructed for natural numbers of the form `bit b n`, - they can be constructed for any given natural number. -/ -@[inline] -def bitCasesOn {motive : Nat → Sort u} (n) (h : ∀ b n, motive (bit b n)) : motive n := - -- `1 &&& n != 0` is faster than `n.testBit 0`. This may change when we have faster `testBit`. - let x := h (1 &&& n != 0) (n >>> 1) - -- `congrArg motive _ ▸ x` is defeq to `x` in non-dependent case - congrArg motive n.bit_testBit_zero_shiftRight_one ▸ x - lemma bit_zero : bit false 0 = 0 := rfl @@ -172,19 +146,6 @@ lemma binaryRec_decreasing (h : n ≠ 0) : div2 n < n := by (lt_of_le_of_ne n.zero_le h.symm) rwa [Nat.mul_one] at this -/-- A recursion principle for `bit` representations of natural numbers. - For a predicate `motive : Nat → Sort*`, if instances can be - constructed for natural numbers of the form `bit b n`, - they can be constructed for all natural numbers. -/ -@[elab_as_elim, specialize] -def binaryRec {motive : Nat → Sort u} (z : motive 0) (f : ∀ b n, motive n → motive (bit b n)) - (n : Nat) : motive n := - if n0 : n = 0 then congrArg motive n0 ▸ z - else - let x := f (1 &&& n != 0) (n >>> 1) (binaryRec z f (n >>> 1)) - congrArg motive n.bit_testBit_zero_shiftRight_one ▸ x -decreasing_by exact bitwise_rec_lemma n0 - /-- `size n` : Returns the size of a natural number in bits i.e. the length of its binary representation -/ def size : ℕ → ℕ := @@ -201,20 +162,6 @@ def bits : ℕ → List Bool := def ldiff : ℕ → ℕ → ℕ := bitwise fun a b => a && not b -@[simp] -lemma binaryRec_zero {motive : Nat → Sort u} - (z : motive 0) (f : ∀ b n, motive n → motive (bit b n)) : - binaryRec z f 0 = z := by - rw [binaryRec] - rfl - -@[simp] -lemma binaryRec_one {C : Nat → Sort u} (z : C 0) (f : ∀ b n, C n → C (bit b n)) : - binaryRec (motive := C) z f 1 = f true 0 z := by - rw [binaryRec] - simp only [succ_ne_self, ↓reduceDIte, reduceShiftRight, binaryRec_zero] - rfl - /-! bitwise ops -/ lemma bodd_bit (b n) : bodd (bit b n) = b := by @@ -242,13 +189,6 @@ lemma shiftLeft'_sub (b m) : ∀ {n k}, k ≤ n → shiftLeft' b m (n - k) = (sh lemma shiftLeft_sub : ∀ (m : Nat) {n k}, k ≤ n → m <<< (n - k) = (m <<< n) >>> k := fun _ _ _ hk => by simp only [← shiftLeft'_false, shiftLeft'_sub false _ hk] --- Not a `simp` lemma, as later `simp` will be able to prove this. -lemma testBit_bit_zero (b n) : testBit (bit b n) 0 = b := by - rw [testBit, bit] - cases b - · simp [← Nat.mul_two] - · simp [← Nat.mul_two] - lemma bodd_eq_one_and_ne_zero : ∀ n, bodd n = (1 &&& n != 0) | 0 => rfl | 1 => rfl @@ -289,24 +229,6 @@ theorem bit_add' : ∀ (b : Bool) (n m : ℕ), bit b (n + m) = bit b n + bit fal theorem bit_ne_zero (b) {n} (h : n ≠ 0) : bit b n ≠ 0 := by cases b <;> dsimp [bit] <;> omega -@[simp] -theorem bit_div_two (b n) : bit b n / 2 = n := by - rw [bit_val, Nat.add_comm, add_mul_div_left, div_eq_of_lt, Nat.zero_add] - · cases b <;> decide - · decide - -@[simp] -theorem bit_shiftRight_one (b n) : bit b n >>> 1 = n := - bit_div_two b n - -@[simp] -theorem bitCasesOn_bit {motive : ℕ → Sort u} (h : ∀ b n, motive (bit b n)) (b : Bool) (n : ℕ) : - bitCasesOn (bit b n) h = h b n := by - change congrArg motive (bit b n).bit_testBit_zero_shiftRight_one ▸ h _ _ = h b n - generalize congrArg motive (bit b n).bit_testBit_zero_shiftRight_one = e; revert e - rw [testBit_bit_zero, bit_shiftRight_one] - intros; rfl - @[simp] theorem bitCasesOn_bit0 {motive : ℕ → Sort u} (H : ∀ b n, motive (bit b n)) (n : ℕ) : bitCasesOn (2 * n) H = H false n := @@ -328,13 +250,6 @@ theorem bit_cases_on_inj {motive : ℕ → Sort u} (H₁ H₂ : ∀ b n, motive ((fun n => bitCasesOn n H₁) = fun n => bitCasesOn n H₂) ↔ H₁ = H₂ := bit_cases_on_injective.eq_iff -@[simp] -theorem bit_eq_zero_iff {n : ℕ} {b : Bool} : bit b n = 0 ↔ n = 0 ∧ b = false := by - constructor - · cases b <;> simp [Nat.bit]; omega - · rintro ⟨rfl, rfl⟩ - rfl - lemma bit_le : ∀ (b : Bool) {m n : ℕ}, m ≤ n → bit b m ≤ bit b n | true, _, _, h => by dsimp [bit]; omega | false, _, _, h => by dsimp [bit]; omega @@ -343,59 +258,6 @@ lemma bit_lt_bit (a b) (h : m < n) : bit a m < bit b n := calc bit a m < 2 * n := by cases a <;> dsimp [bit] <;> omega _ ≤ bit b n := by cases b <;> dsimp [bit] <;> omega -/-- -The same as `binaryRec_eq`, -but that one unfortunately requires `f` to be the identity when appending `false` to `0`. -Here, we allow you to explicitly say that that case is not happening, -i.e. supplying `n = 0 → b = true`. -/ -theorem binaryRec_eq' {motive : ℕ → Sort*} {z : motive 0} {f : ∀ b n, motive n → motive (bit b n)} - (b n) (h : f false 0 z = z ∨ (n = 0 → b = true)) : - binaryRec z f (bit b n) = f b n (binaryRec z f n) := by - by_cases h' : bit b n = 0 - case pos => - obtain ⟨rfl, rfl⟩ := bit_eq_zero_iff.mp h' - simp only [imp_false, or_false, eq_self_iff_true, not_true, reduceCtorEq] at h - unfold binaryRec - exact h.symm - case neg => - rw [binaryRec, dif_neg h'] - change congrArg motive (bit b n).bit_testBit_zero_shiftRight_one ▸ f _ _ _ = _ - generalize congrArg motive (bit b n).bit_testBit_zero_shiftRight_one = e; revert e - rw [testBit_bit_zero, bit_shiftRight_one] - intros; rfl - -theorem binaryRec_eq {motive : Nat → Sort u} {z : motive 0} - {f : ∀ b n, motive n → motive (bit b n)} - (h : f false 0 z = z) (b n) : - binaryRec z f (bit b n) = f b n (binaryRec z f n) := - binaryRec_eq' b n (.inl h) - -/-- The same as `binaryRec`, but the induction step can assume that if `n=0`, - the bit being appended is `true`-/ -@[elab_as_elim, specialize] -def binaryRec' {motive : ℕ → Sort*} (z : motive 0) - (f : ∀ b n, (n = 0 → b = true) → motive n → motive (bit b n)) : - ∀ n, motive n := - binaryRec z fun b n ih => - if h : n = 0 → b = true then f b n h ih - else - have : bit b n = 0 := by - rw [bit_eq_zero_iff] - cases n <;> cases b <;> simp at h ⊢ - congrArg motive this ▸ z - -/-- The same as `binaryRec`, but special casing both 0 and 1 as base cases -/ -@[elab_as_elim, specialize] -def binaryRecFromOne {motive : ℕ → Sort*} (z₀ : motive 0) (z₁ : motive 1) - (f : ∀ b n, n ≠ 0 → motive n → motive (bit b n)) : - ∀ n, motive n := - binaryRec' z₀ fun b n h ih => - if h' : n = 0 then - have : bit b n = bit true 0 := by - rw [h', h h'] - congrArg motive this ▸ z₁ - else f b n h' ih - @[simp] theorem zero_bits : bits 0 = [] := by simp [Nat.bits] diff --git a/Mathlib/Data/Nat/Bitwise.lean b/Mathlib/Data/Nat/Bitwise.lean index 07873ec6fd976..c0cbec11309d7 100644 --- a/Mathlib/Data/Nat/Bitwise.lean +++ b/Mathlib/Data/Nat/Bitwise.lean @@ -3,13 +3,14 @@ Copyright (c) 2020 Markus Himmel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Markus Himmel, Alex Keizer -/ +import Mathlib.Algebra.Group.Units.Basic +import Mathlib.Algebra.NeZero +import Mathlib.Algebra.Ring.Nat import Mathlib.Data.List.GetD import Mathlib.Data.Nat.Bits -import Mathlib.Algebra.Ring.Nat import Mathlib.Order.Basic import Mathlib.Tactic.AdaptationNote import Mathlib.Tactic.Common -import Mathlib.Algebra.NeZero /-! # Bitwise operations on natural numbers @@ -88,11 +89,6 @@ lemma bitwise_bit {f : Bool → Bool → Bool} (h : f false false = false := by cases a <;> cases b <;> simp [h2, h4] <;> split_ifs <;> simp_all (config := {decide := true}) [two_mul] -@[simp] -lemma bit_mod_two (a : Bool) (x : ℕ) : - bit a x % 2 = a.toNat := by - cases a <;> simp [bit_val, mul_add_mod] - lemma bit_mod_two_eq_zero_iff (a x) : bit a x % 2 = 0 ↔ !a := by simp diff --git a/Mathlib/Data/Nat/EvenOddRec.lean b/Mathlib/Data/Nat/EvenOddRec.lean index a5c452c173847..ce6c89abcd14f 100644 --- a/Mathlib/Data/Nat/EvenOddRec.lean +++ b/Mathlib/Data/Nat/EvenOddRec.lean @@ -4,7 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Stuart Presnell -/ import Mathlib.Algebra.Ring.Parity -import Mathlib.Data.Nat.Bits +import Mathlib.Data.Nat.BinaryRec + /-! # A recursion principle based on even and odd numbers. -/ namespace Nat diff --git a/Mathlib/Data/Nat/Fib/Basic.lean b/Mathlib/Data/Nat/Fib/Basic.lean index 587e609b5ab7b..399cabbe6107f 100644 --- a/Mathlib/Data/Nat/Fib/Basic.lean +++ b/Mathlib/Data/Nat/Fib/Basic.lean @@ -5,8 +5,9 @@ Authors: Kevin Kappelmann, Kyle Miller, Mario Carneiro -/ import Mathlib.Algebra.BigOperators.Group.Finset import Mathlib.Data.Finset.NatAntidiagonal -import Mathlib.Data.Nat.Bits import Mathlib.Data.Nat.GCD.Basic +import Mathlib.Data.Nat.BinaryRec +import Mathlib.Logic.Function.Iterate import Mathlib.Tactic.Ring import Mathlib.Tactic.Zify diff --git a/Mathlib/Data/Num/Basic.lean b/Mathlib/Data/Num/Basic.lean index 5e35ca5550ad8..07f592d41ee94 100644 --- a/Mathlib/Data/Num/Basic.lean +++ b/Mathlib/Data/Num/Basic.lean @@ -4,9 +4,11 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Mario Carneiro -/ import Lean.Linter.Deprecated -import Mathlib.Data.Int.Notation import Mathlib.Algebra.Group.ZeroOne -import Mathlib.Data.Nat.Bits +import Mathlib.Data.Int.Notation +import Mathlib.Data.Nat.BinaryRec +import Mathlib.Tactic.TypeStar + /-! # Binary representation of integers using inductive types diff --git a/Mathlib/Data/Tree/Get.lean b/Mathlib/Data/Tree/Get.lean index e3adeba9eea04..b374b675b07ba 100644 --- a/Mathlib/Data/Tree/Get.lean +++ b/Mathlib/Data/Tree/Get.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Wojciech Nawrocki -/ import Mathlib.Data.Num.Basic +import Mathlib.Data.Ordering.Basic import Mathlib.Data.Tree.Basic /-! diff --git a/scripts/noshake.json b/scripts/noshake.json index a98fca062e875..1fa2973a3a767 100644 --- a/scripts/noshake.json +++ b/scripts/noshake.json @@ -41,6 +41,7 @@ "Mathlib.Control.Traversable.Lemmas", "Mathlib.Data.Finsupp.Notation", "Mathlib.Data.Int.Defs", + "Mathlib.Data.Int.Notation", "Mathlib.Data.Matrix.Notation", "Mathlib.Data.Matroid.Init", "Mathlib.Data.Nat.Notation", @@ -336,7 +337,6 @@ "Mathlib.Logic.Function.Defs": ["Mathlib.Tactic.AdaptationNote"], "Mathlib.Logic.Function.Basic": ["Batteries.Tactic.Init"], "Mathlib.Logic.Equiv.Defs": ["Mathlib.Data.Bool.Basic"], - "Mathlib.Logic.Basic": ["Mathlib.Data.Int.Notation"], "Mathlib.LinearAlgebra.Matrix.Transvection": ["Mathlib.Data.Matrix.DMatrix"], "Mathlib.LinearAlgebra.Basic": ["Mathlib.Algebra.BigOperators.Group.Finset"], "Mathlib.LinearAlgebra.AffineSpace.FiniteDimensional": ["Mathlib.Init.Core"], @@ -370,7 +370,6 @@ "Mathlib.Data.List.Basic": ["Mathlib.Control.Basic", "Mathlib.Data.Option.Basic"], "Mathlib.Data.LazyList.Basic": ["Mathlib.Lean.Thunk"], - "Mathlib.Data.Int.Order.Basic": ["Mathlib.Data.Int.Notation"], "Mathlib.Data.Int.Defs": ["Batteries.Data.Int.Order"], "Mathlib.Data.FunLike.Basic": ["Mathlib.Logic.Function.Basic"], "Mathlib.Data.Finset.Basic": ["Mathlib.Data.Finset.Attr"], @@ -423,7 +422,6 @@ "Mathlib.Algebra.Group.Units.Basic": ["Mathlib.Tactic.Subsingleton"], "Mathlib.Algebra.Group.Units": ["Mathlib.Tactic.Subsingleton"], "Mathlib.Algebra.Group.Pi.Basic": ["Batteries.Tactic.Classical"], - "Mathlib.Algebra.Group.Defs": ["Mathlib.Data.Int.Notation"], "Mathlib.Algebra.GeomSum": ["Mathlib.Algebra.Order.BigOperators.Ring.Finset"], "Mathlib.Algebra.Category.Ring.Basic": ["Mathlib.CategoryTheory.ConcreteCategory.ReflectsIso"], From d172902852dabd02b61ebafc692e751d29ecac1e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 20 Oct 2024 12:53:16 +0000 Subject: [PATCH 201/505] chore(Algebra/Order/BigOperators/Group/Finset): don't import `Ring` (#17939) By moving the positivity extension, one can shave off a bunch of imports --- Mathlib/Algebra/BigOperators/Finprod.lean | 1 + Mathlib/Algebra/Order/Antidiag/Finsupp.lean | 1 - .../Order/BigOperators/Group/Finset.lean | 45 ++----------------- Mathlib/Algebra/Order/Chebyshev.lean | 1 - Mathlib/Algebra/Order/Interval/Basic.lean | 1 + Mathlib/Analysis/Analytic/Inverse.lean | 1 + Mathlib/Analysis/Complex/AbelLimit.lean | 1 + .../Enumerative/Composition.lean | 3 +- .../Enumerative/DoubleCounting.lean | 1 + Mathlib/Combinatorics/Pigeonhole.lean | 4 +- .../SetFamily/AhlswedeZhang.lean | 2 - .../SetFamily/FourFunctions.lean | 1 - .../SimpleGraph/Regularity/Energy.lean | 1 - Mathlib/Data/Finsupp/Weight.lean | 4 -- Mathlib/Data/Nat/Choose/Sum.lean | 1 - Mathlib/Data/Set/Equitable.lean | 4 +- Mathlib/Data/Sign.lean | 2 + Mathlib/LinearAlgebra/Matrix/DotProduct.lean | 3 -- Mathlib/LinearAlgebra/Multilinear/Basic.lean | 2 - Mathlib/NumberTheory/Divisors.lean | 2 +- Mathlib/NumberTheory/Harmonic/Defs.lean | 1 + Mathlib/Order/Partition/Equipartition.lean | 1 + Mathlib/Tactic/Positivity/Finset.lean | 37 +++++++++++++++ Mathlib/Testing/SlimCheck/Functions.lean | 2 +- Mathlib/Topology/Algebra/GroupWithZero.lean | 1 + .../Topology/Algebra/WithZeroTopology.lean | 1 + scripts/noshake.json | 3 +- test/positivity.lean | 1 + 28 files changed, 58 insertions(+), 70 deletions(-) diff --git a/Mathlib/Algebra/BigOperators/Finprod.lean b/Mathlib/Algebra/BigOperators/Finprod.lean index b6b2d96428073..33329a6e81269 100644 --- a/Mathlib/Algebra/BigOperators/Finprod.lean +++ b/Mathlib/Algebra/BigOperators/Finprod.lean @@ -7,6 +7,7 @@ import Mathlib.Algebra.BigOperators.GroupWithZero.Finset import Mathlib.Algebra.Group.FiniteSupport import Mathlib.Algebra.NoZeroSMulDivisors.Basic import Mathlib.Algebra.Order.BigOperators.Group.Finset +import Mathlib.Algebra.Order.Ring.Defs import Mathlib.Data.Set.Subsingleton /-! diff --git a/Mathlib/Algebra/Order/Antidiag/Finsupp.lean b/Mathlib/Algebra/Order/Antidiag/Finsupp.lean index 3564d4ffb4076..150ff45ab21d7 100644 --- a/Mathlib/Algebra/Order/Antidiag/Finsupp.lean +++ b/Mathlib/Algebra/Order/Antidiag/Finsupp.lean @@ -5,7 +5,6 @@ Authors: Antoine Chambert-Loir, María Inés de Frutos-Fernández, Eric Wieser, Yaël Dillies -/ import Mathlib.Algebra.Order.Antidiag.Pi -import Mathlib.Algebra.Order.BigOperators.Group.Finset import Mathlib.Data.Finsupp.Basic /-! diff --git a/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean b/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean index 8593406fe5e67..f799eabddb523 100644 --- a/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean +++ b/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean @@ -5,10 +5,9 @@ Authors: Johannes Hölzl -/ import Mathlib.Algebra.BigOperators.Group.Finset import Mathlib.Algebra.Order.BigOperators.Group.Multiset +import Mathlib.Algebra.Order.Group.Nat import Mathlib.Data.Multiset.OrderedMonoid import Mathlib.Tactic.Bound.Attribute -import Mathlib.Tactic.NormNum.Basic -import Mathlib.Tactic.Positivity.Core /-! # Big operators on a finset in ordered groups @@ -17,6 +16,8 @@ This file contains the results concerning the interaction of multiset big operat groups/monoids. -/ +assert_not_exists Ring + open Function variable {ι α β M N G k R : Type*} @@ -581,43 +582,3 @@ theorem sup_powerset_len {α : Type*} [DecidableEq α] (x : Multiset α) : Eq.symm (finset_sum_eq_sup_iff_disjoint.mpr fun _ _ _ _ h => pairwise_disjoint_powersetCard x h) end Multiset - -namespace Mathlib.Meta.Positivity -open Qq Lean Meta Finset - -/-- The `positivity` extension which proves that `∑ i ∈ s, f i` is nonnegative if `f` is, and -positive if each `f i` is and `s` is nonempty. - -TODO: The following example does not work -``` -example (s : Finset ℕ) (f : ℕ → ℤ) (hf : ∀ n, 0 ≤ f n) : 0 ≤ s.sum f := by positivity -``` -because `compareHyp` can't look for assumptions behind binders. --/ -@[positivity Finset.sum _ _] -def evalFinsetSum : PositivityExt where eval {u α} zα pα e := do - match e with - | ~q(@Finset.sum $ι _ $instα $s $f) => - let i : Q($ι) ← mkFreshExprMVarQ q($ι) .syntheticOpaque - have body : Q($α) := .betaRev f #[i] - let rbody ← core zα pα body - let p_pos : Option Q(0 < $e) := ← (do - let .positive pbody := rbody | pure none -- Fail if the body is not provably positive - let .some ps ← proveFinsetNonempty s | pure none - let .some pα' ← trySynthInstanceQ q(OrderedCancelAddCommMonoid $α) | pure none - assertInstancesCommute - let pr : Q(∀ i, 0 < $f i) ← mkLambdaFVars #[i] pbody - return some q(@sum_pos $ι $α $pα' $f $s (fun i _ ↦ $pr i) $ps)) - -- Try to show that the sum is positive - if let some p_pos := p_pos then - return .positive p_pos - -- Fall back to showing that the sum is nonnegative - else - let pbody ← rbody.toNonneg - let pr : Q(∀ i, 0 ≤ $f i) ← mkLambdaFVars #[i] pbody - let pα' ← synthInstanceQ q(OrderedAddCommMonoid $α) - assertInstancesCommute - return .nonnegative q(@sum_nonneg $ι $α $pα' $f $s fun i _ ↦ $pr i) - | _ => throwError "not Finset.sum" - -end Mathlib.Meta.Positivity diff --git a/Mathlib/Algebra/Order/Chebyshev.lean b/Mathlib/Algebra/Order/Chebyshev.lean index 439d79b0c4e2b..7987886b72b00 100644 --- a/Mathlib/Algebra/Order/Chebyshev.lean +++ b/Mathlib/Algebra/Order/Chebyshev.lean @@ -3,7 +3,6 @@ Copyright (c) 2023 Mantas Bakšys, Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mantas Bakšys, Yaël Dillies -/ -import Mathlib.Algebra.Order.BigOperators.Group.Finset import Mathlib.Algebra.Order.Monovary import Mathlib.Algebra.Order.Rearrangement import Mathlib.GroupTheory.Perm.Cycle.Basic diff --git a/Mathlib/Algebra/Order/Interval/Basic.lean b/Mathlib/Algebra/Order/Interval/Basic.lean index 5fc5f067acadf..bcfec00e2cd85 100644 --- a/Mathlib/Algebra/Order/Interval/Basic.lean +++ b/Mathlib/Algebra/Order/Interval/Basic.lean @@ -6,6 +6,7 @@ Authors: Yaël Dillies import Mathlib.Algebra.Group.Pointwise.Set.Basic import Mathlib.Algebra.Order.BigOperators.Group.Finset import Mathlib.Order.Interval.Basic +import Mathlib.Tactic.Positivity.Core /-! # Interval arithmetic diff --git a/Mathlib/Analysis/Analytic/Inverse.lean b/Mathlib/Analysis/Analytic/Inverse.lean index 34ad1130c225c..4b94111adff37 100644 --- a/Mathlib/Analysis/Analytic/Inverse.lean +++ b/Mathlib/Analysis/Analytic/Inverse.lean @@ -5,6 +5,7 @@ Authors: Sébastien Gouëzel -/ import Mathlib.Analysis.Analytic.Composition import Mathlib.Analysis.Analytic.Linear +import Mathlib.Tactic.Positivity.Finset /-! diff --git a/Mathlib/Analysis/Complex/AbelLimit.lean b/Mathlib/Analysis/Complex/AbelLimit.lean index 20c8a9420d6f3..25d7f649514c6 100644 --- a/Mathlib/Analysis/Complex/AbelLimit.lean +++ b/Mathlib/Analysis/Complex/AbelLimit.lean @@ -6,6 +6,7 @@ Authors: Jeremy Tan import Mathlib.Analysis.Complex.Basic import Mathlib.Analysis.SpecificLimits.Normed import Mathlib.Tactic.Peel +import Mathlib.Tactic.Positivity.Finset /-! # Abel's limit theorem diff --git a/Mathlib/Combinatorics/Enumerative/Composition.lean b/Mathlib/Combinatorics/Enumerative/Composition.lean index 917e4a7e41dbc..c825e38016f80 100644 --- a/Mathlib/Combinatorics/Enumerative/Composition.lean +++ b/Mathlib/Combinatorics/Enumerative/Composition.lean @@ -6,7 +6,6 @@ Authors: Sébastien Gouëzel import Mathlib.Algebra.BigOperators.Fin import Mathlib.Algebra.Order.BigOperators.Group.Finset import Mathlib.Data.Finset.Sort -import Mathlib.Data.Set.Subsingleton /-! # Compositions @@ -807,7 +806,7 @@ theorem card_boundaries_eq_succ_length : c.boundaries.card = c.length + 1 := theorem length_lt_card_boundaries : c.length < c.boundaries.card := by rw [c.card_boundaries_eq_succ_length] - exact lt_add_one _ + exact Nat.lt_add_one _ theorem lt_length (i : Fin c.length) : (i : ℕ) + 1 < c.boundaries.card := lt_tsub_iff_right.mp i.2 diff --git a/Mathlib/Combinatorics/Enumerative/DoubleCounting.lean b/Mathlib/Combinatorics/Enumerative/DoubleCounting.lean index b929f72e6e626..3061124939269 100644 --- a/Mathlib/Combinatorics/Enumerative/DoubleCounting.lean +++ b/Mathlib/Combinatorics/Enumerative/DoubleCounting.lean @@ -5,6 +5,7 @@ Authors: Yaël Dillies -/ import Mathlib.Algebra.BigOperators.Ring import Mathlib.Algebra.Order.BigOperators.Group.Finset +import Mathlib.Algebra.Order.Ring.Nat /-! # Double countings diff --git a/Mathlib/Combinatorics/Pigeonhole.lean b/Mathlib/Combinatorics/Pigeonhole.lean index 3811b73e8b204..b91d6c658fd57 100644 --- a/Mathlib/Combinatorics/Pigeonhole.lean +++ b/Mathlib/Combinatorics/Pigeonhole.lean @@ -4,10 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kyle Miller, Yury Kudryashov -/ import Mathlib.Algebra.Module.BigOperators -import Mathlib.Algebra.Module.Defs -import Mathlib.Algebra.Order.BigOperators.Group.Finset +import Mathlib.Algebra.Order.Ring.Nat import Mathlib.Data.Nat.ModEq -import Mathlib.Data.Set.Finite /-! # Pigeonhole principles diff --git a/Mathlib/Combinatorics/SetFamily/AhlswedeZhang.lean b/Mathlib/Combinatorics/SetFamily/AhlswedeZhang.lean index ea08e4e2baeef..73740d7fb77d3 100644 --- a/Mathlib/Combinatorics/SetFamily/AhlswedeZhang.lean +++ b/Mathlib/Combinatorics/SetFamily/AhlswedeZhang.lean @@ -6,10 +6,8 @@ Authors: Yaël Dillies, Vladimir Ivanov import Mathlib.Algebra.BigOperators.Intervals import Mathlib.Algebra.BigOperators.Ring import Mathlib.Algebra.Order.BigOperators.Group.Finset -import Mathlib.Algebra.Order.Field.Basic import Mathlib.Data.Finset.Sups import Mathlib.Tactic.FieldSimp -import Mathlib.Tactic.Positivity.Basic import Mathlib.Tactic.Ring /-! diff --git a/Mathlib/Combinatorics/SetFamily/FourFunctions.lean b/Mathlib/Combinatorics/SetFamily/FourFunctions.lean index d3bd37943cd19..df03a09a30705 100644 --- a/Mathlib/Combinatorics/SetFamily/FourFunctions.lean +++ b/Mathlib/Combinatorics/SetFamily/FourFunctions.lean @@ -7,7 +7,6 @@ import Mathlib.Algebra.Order.BigOperators.Group.Finset import Mathlib.Algebra.Order.Pi import Mathlib.Algebra.Order.Ring.Basic import Mathlib.Data.Finset.Sups -import Mathlib.Data.Set.Subsingleton import Mathlib.Order.Birkhoff import Mathlib.Order.Booleanisation import Mathlib.Order.Sublattice diff --git a/Mathlib/Combinatorics/SimpleGraph/Regularity/Energy.lean b/Mathlib/Combinatorics/SimpleGraph/Regularity/Energy.lean index 54e5c27b5c906..ad2bb21be97a3 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Regularity/Energy.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Regularity/Energy.lean @@ -5,7 +5,6 @@ Authors: Yaël Dillies, Bhavik Mehta -/ import Mathlib.Algebra.Module.Defs import Mathlib.Algebra.Order.BigOperators.Group.Finset -import Mathlib.Algebra.Order.Ring.Basic import Mathlib.Combinatorics.SimpleGraph.Density import Mathlib.Data.Rat.BigOperators diff --git a/Mathlib/Data/Finsupp/Weight.lean b/Mathlib/Data/Finsupp/Weight.lean index ed767ee615ac8..075821985671b 100644 --- a/Mathlib/Data/Finsupp/Weight.lean +++ b/Mathlib/Data/Finsupp/Weight.lean @@ -3,11 +3,7 @@ Copyright (c) 2024 Antoine Chambert-Loir, María Inés de Frutos-Fernández. All Released under Apache 2.0 license as described in the file LICENSE. Authors: Antoine Chambert-Loir, María Inés de Frutos-Fernández -/ - -import Mathlib.Algebra.Order.BigOperators.Group.Finset import Mathlib.Algebra.Order.Module.Defs -import Mathlib.Algebra.Order.Ring.Defs -import Mathlib.Algebra.Order.Monoid.Canonical.Defs import Mathlib.LinearAlgebra.Finsupp /-! # weights of Finsupp functions diff --git a/Mathlib/Data/Nat/Choose/Sum.lean b/Mathlib/Data/Nat/Choose/Sum.lean index 04851939e6336..ab6812bc6fcb7 100644 --- a/Mathlib/Data/Nat/Choose/Sum.lean +++ b/Mathlib/Data/Nat/Choose/Sum.lean @@ -7,7 +7,6 @@ import Mathlib.Algebra.BigOperators.Intervals import Mathlib.Algebra.BigOperators.NatAntidiagonal import Mathlib.Algebra.BigOperators.Ring import Mathlib.Algebra.Order.BigOperators.Group.Finset -import Mathlib.Data.Nat.Choose.Basic import Mathlib.Tactic.Linarith import Mathlib.Tactic.Ring diff --git a/Mathlib/Data/Set/Equitable.lean b/Mathlib/Data/Set/Equitable.lean index 8eb08216d19ac..b2e7669c813a3 100644 --- a/Mathlib/Data/Set/Equitable.lean +++ b/Mathlib/Data/Set/Equitable.lean @@ -3,10 +3,8 @@ Copyright (c) 2021 Yaël Dillies, Bhavik Mehta. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies, Bhavik Mehta -/ -import Mathlib.Data.Set.Subsingleton import Mathlib.Algebra.Order.BigOperators.Group.Finset -import Mathlib.Algebra.Group.Nat -import Mathlib.Data.Set.Basic +import Mathlib.Algebra.Order.Ring.Defs /-! # Equitable functions diff --git a/Mathlib/Data/Sign.lean b/Mathlib/Data/Sign.lean index 7e0033c7d8a2a..a492a43728d98 100644 --- a/Mathlib/Data/Sign.lean +++ b/Mathlib/Data/Sign.lean @@ -5,7 +5,9 @@ Authors: Eric Rodriguez -/ import Mathlib.Algebra.GroupWithZero.Units.Lemmas import Mathlib.Algebra.Order.BigOperators.Group.Finset +import Mathlib.Algebra.Order.Ring.Cast import Mathlib.Data.Fintype.BigOperators + /-! # Sign function diff --git a/Mathlib/LinearAlgebra/Matrix/DotProduct.lean b/Mathlib/LinearAlgebra/Matrix/DotProduct.lean index bb06aeea3eee8..bf76ac67e8efa 100644 --- a/Mathlib/LinearAlgebra/Matrix/DotProduct.lean +++ b/Mathlib/LinearAlgebra/Matrix/DotProduct.lean @@ -3,11 +3,8 @@ Copyright (c) 2019 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Patrick Massot, Casper Putz, Anne Baanen -/ -import Mathlib.Algebra.Ring.Regular import Mathlib.Algebra.Order.Star.Basic -import Mathlib.Data.Matrix.Basic import Mathlib.LinearAlgebra.StdBasis -import Mathlib.Algebra.Order.BigOperators.Group.Finset /-! # Dot product of two vectors diff --git a/Mathlib/LinearAlgebra/Multilinear/Basic.lean b/Mathlib/LinearAlgebra/Multilinear/Basic.lean index ea1242276149d..f37192686c399 100644 --- a/Mathlib/LinearAlgebra/Multilinear/Basic.lean +++ b/Mathlib/LinearAlgebra/Multilinear/Basic.lean @@ -5,10 +5,8 @@ Authors: Sébastien Gouëzel -/ import Mathlib.Algebra.Algebra.Defs import Mathlib.Algebra.NoZeroSMulDivisors.Pi -import Mathlib.Algebra.Order.BigOperators.Group.Finset import Mathlib.Data.Fintype.BigOperators import Mathlib.Data.Fintype.Sort -import Mathlib.Data.List.FinRange import Mathlib.LinearAlgebra.Pi import Mathlib.Logic.Equiv.Fintype import Mathlib.Tactic.Abel diff --git a/Mathlib/NumberTheory/Divisors.lean b/Mathlib/NumberTheory/Divisors.lean index 3fb893b14cd82..d0f8d0cc61130 100644 --- a/Mathlib/NumberTheory/Divisors.lean +++ b/Mathlib/NumberTheory/Divisors.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Aaron Anderson -/ import Mathlib.Algebra.Order.BigOperators.Group.Finset -import Mathlib.Data.Nat.Prime.Basic +import Mathlib.Algebra.Order.Ring.Nat import Mathlib.Data.Nat.PrimeFin import Mathlib.Order.Interval.Finset.Nat diff --git a/Mathlib/NumberTheory/Harmonic/Defs.lean b/Mathlib/NumberTheory/Harmonic/Defs.lean index a7887fff3ff84..1169d86a0e470 100644 --- a/Mathlib/NumberTheory/Harmonic/Defs.lean +++ b/Mathlib/NumberTheory/Harmonic/Defs.lean @@ -7,6 +7,7 @@ import Mathlib.Algebra.BigOperators.Intervals import Mathlib.Algebra.Order.BigOperators.Ring.Finset import Mathlib.Algebra.Order.Field.Basic import Mathlib.Tactic.Linarith +import Mathlib.Tactic.Positivity.Finset /-! diff --git a/Mathlib/Order/Partition/Equipartition.lean b/Mathlib/Order/Partition/Equipartition.lean index 7833e98aa3fe8..3969b84ca5eb0 100644 --- a/Mathlib/Order/Partition/Equipartition.lean +++ b/Mathlib/Order/Partition/Equipartition.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Yaël Dillies, Bhavik Mehta. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies, Bhavik Mehta -/ +import Mathlib.Algebra.Order.Ring.Nat import Mathlib.Data.Set.Equitable import Mathlib.Logic.Equiv.Fin import Mathlib.Order.Partition.Finpartition diff --git a/Mathlib/Tactic/Positivity/Finset.lean b/Mathlib/Tactic/Positivity/Finset.lean index 6326f5d927658..bef478ebd38c6 100644 --- a/Mathlib/Tactic/Positivity/Finset.lean +++ b/Mathlib/Tactic/Positivity/Finset.lean @@ -3,7 +3,9 @@ Copyright (c) 2023 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ +import Mathlib.Algebra.Order.BigOperators.Group.Finset import Mathlib.Data.Finset.Density +import Mathlib.Tactic.NormNum.Basic import Mathlib.Tactic.Positivity.Core /-! @@ -52,6 +54,41 @@ def evalFinsetDens : PositivityExt where eval {u 𝕜} _ _ e := do return .positive q(@Nonempty.dens_pos $α $instα $s $ps) | _, _, _ => throwError "not Finset.dens" +/-- The `positivity` extension which proves that `∑ i ∈ s, f i` is nonnegative if `f` is, and +positive if each `f i` is and `s` is nonempty. + +TODO: The following example does not work +``` +example (s : Finset ℕ) (f : ℕ → ℤ) (hf : ∀ n, 0 ≤ f n) : 0 ≤ s.sum f := by positivity +``` +because `compareHyp` can't look for assumptions behind binders. +-/ +@[positivity Finset.sum _ _] +def evalFinsetSum : PositivityExt where eval {u α} zα pα e := do + match e with + | ~q(@Finset.sum $ι _ $instα $s $f) => + let i : Q($ι) ← mkFreshExprMVarQ q($ι) .syntheticOpaque + have body : Q($α) := .betaRev f #[i] + let rbody ← core zα pα body + let p_pos : Option Q(0 < $e) := ← (do + let .positive pbody := rbody | pure none -- Fail if the body is not provably positive + let .some ps ← proveFinsetNonempty s | pure none + let .some pα' ← trySynthInstanceQ q(OrderedCancelAddCommMonoid $α) | pure none + assertInstancesCommute + let pr : Q(∀ i, 0 < $f i) ← mkLambdaFVars #[i] pbody + return some q(@sum_pos $ι $α $pα' $f $s (fun i _ ↦ $pr i) $ps)) + -- Try to show that the sum is positive + if let some p_pos := p_pos then + return .positive p_pos + -- Fall back to showing that the sum is nonnegative + else + let pbody ← rbody.toNonneg + let pr : Q(∀ i, 0 ≤ $f i) ← mkLambdaFVars #[i] pbody + let pα' ← synthInstanceQ q(OrderedAddCommMonoid $α) + assertInstancesCommute + return .nonnegative q(@sum_nonneg $ι $α $pα' $f $s fun i _ ↦ $pr i) + | _ => throwError "not Finset.sum" + variable {α : Type*} {s : Finset α} example : 0 ≤ s.card := by positivity diff --git a/Mathlib/Testing/SlimCheck/Functions.lean b/Mathlib/Testing/SlimCheck/Functions.lean index 51eadae43cdec..7b73daf4bdbcb 100644 --- a/Mathlib/Testing/SlimCheck/Functions.lean +++ b/Mathlib/Testing/SlimCheck/Functions.lean @@ -289,7 +289,7 @@ theorem List.applyId_zip_eq [DecidableEq α] {xs ys : List α} (h₀ : List.Nodu | nil => cases h₂ | cons x' xs xs_ih => cases i - · simp only [length_cons, lt_add_iff_pos_left, add_pos_iff, zero_lt_one, or_true, + · simp only [length_cons, lt_add_iff_pos_left, add_pos_iff, Nat.lt_add_one, or_true, getElem?_eq_getElem, getElem_cons_zero, Option.some.injEq] at h₂ subst h₂ cases ys diff --git a/Mathlib/Topology/Algebra/GroupWithZero.lean b/Mathlib/Topology/Algebra/GroupWithZero.lean index b601051b87dc5..8874a996043c3 100644 --- a/Mathlib/Topology/Algebra/GroupWithZero.lean +++ b/Mathlib/Topology/Algebra/GroupWithZero.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import Mathlib.Algebra.Group.Pi.Lemmas +import Mathlib.Algebra.GroupWithZero.Units.Equiv import Mathlib.Topology.Algebra.Monoid import Mathlib.Topology.Homeomorph diff --git a/Mathlib/Topology/Algebra/WithZeroTopology.lean b/Mathlib/Topology/Algebra/WithZeroTopology.lean index 8b290b9995546..59382fdd867a4 100644 --- a/Mathlib/Topology/Algebra/WithZeroTopology.lean +++ b/Mathlib/Topology/Algebra/WithZeroTopology.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Patrick Massot. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Patrick Massot -/ +import Mathlib.Algebra.Order.GroupWithZero.Canonical import Mathlib.Topology.Algebra.GroupWithZero import Mathlib.Topology.Order.OrderClosed diff --git a/scripts/noshake.json b/scripts/noshake.json index 1fa2973a3a767..5ee9b18083934 100644 --- a/scripts/noshake.json +++ b/scripts/noshake.json @@ -219,7 +219,7 @@ "Mathlib.Tactic.ProxyType": ["Mathlib.Logic.Equiv.Defs"], "Mathlib.Tactic.ProdAssoc": ["Mathlib.Logic.Equiv.Defs"], "Mathlib.Tactic.Positivity.Finset": - ["Mathlib.Data.Finset.Density", "Mathlib.Data.Fintype.Card"], + ["Mathlib.Algebra.Order.BigOperators.Group.Finset", "Mathlib.Data.Finset.Density"], "Mathlib.Tactic.Positivity.Basic": ["Mathlib.Algebra.GroupPower.Order", "Mathlib.Algebra.Order.Group.PosPart", @@ -295,7 +295,6 @@ ["Mathlib.Algebra.Group.ZeroOne", "Mathlib.Tactic.Bound.Init"], "Mathlib.Tactic.Bound": ["Mathlib.Algebra.Order.AbsoluteValue", - "Mathlib.Algebra.Order.BigOperators.Group.Finset", "Mathlib.Algebra.Order.Floor", "Mathlib.Analysis.Analytic.Basic", "Mathlib.Tactic.Bound.Init"], diff --git a/test/positivity.lean b/test/positivity.lean index fc773c8abd962..07295214c3bbe 100644 --- a/test/positivity.lean +++ b/test/positivity.lean @@ -4,6 +4,7 @@ import Mathlib.Analysis.Normed.Group.Basic import Mathlib.Analysis.SpecialFunctions.Pow.Real import Mathlib.Analysis.SpecialFunctions.Log.Basic import Mathlib.MeasureTheory.Integral.Bochner +import Mathlib.Tactic.Positivity.Finset /-! # Tests for the `positivity` tactic From 6cd5a568ec89051613dd22e5053bbb64b11b575b Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Sun, 20 Oct 2024 13:19:34 +0000 Subject: [PATCH 202/505] feat: induction principles for ContinuousMap via Stone-Weierstrass (#17831) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR provides induction lemmas for `C(s, 𝕜)` and `C(s, 𝕜)₀` with `RCLike 𝕜`. The idea is that when `s` is compact, the `closure` hypothesis will follow from the Stone-Weierstrass theorem and a limit property of the motive `p`. --- .../ContinuousMap/StoneWeierstrass.lean | 65 ++++++++++++++++++- 1 file changed, 64 insertions(+), 1 deletion(-) diff --git a/Mathlib/Topology/ContinuousMap/StoneWeierstrass.lean b/Mathlib/Topology/ContinuousMap/StoneWeierstrass.lean index 8097bb16a5cf8..ecfa6bfa4e8f4 100644 --- a/Mathlib/Topology/ContinuousMap/StoneWeierstrass.lean +++ b/Mathlib/Topology/ContinuousMap/StoneWeierstrass.lean @@ -421,6 +421,37 @@ theorem polynomialFunctions.starClosure_topologicalClosure {𝕜 : Type*} [RCLik ContinuousMap.starSubalgebra_topologicalClosure_eq_top_of_separatesPoints _ (Subalgebra.separatesPoints_monotone le_sup_left (polynomialFunctions_separatesPoints s)) +/-- An induction principle for `C(s, 𝕜)`. -/ +@[elab_as_elim] +theorem ContinuousMap.induction_on {𝕜 : Type*} [RCLike 𝕜] {s : Set 𝕜} + {p : C(s, 𝕜) → Prop} (const : ∀ r, p (.const s r)) (id : p (.restrict s <| .id 𝕜)) + (add : ∀ f g, p f → p g → p (f + g)) (mul : ∀ f g, p f → p g → p (f * g)) + (star : ∀ f, p f → p (star f)) + (closure : (∀ f ∈ (polynomialFunctions s).starClosure, p f) → ∀ f, p f) (f : C(s, 𝕜)) : + p f := by + refine closure (fun f hf => ?_) f + rw [polynomialFunctions.starClosure_eq_adjoin_X] at hf + induction hf using StarAlgebra.adjoin_induction with + | mem f hf => + simp only [toContinuousMapOnAlgHom_apply, Set.mem_singleton_iff] at hf + rwa [hf, toContinuousMapOn_X_eq_restrict_id] + | algebraMap r => exact const r + | add _ _ _ _ hf hg => exact add _ _ hf hg + | mul _ _ _ _ hf hg => exact mul _ _ hf hg + | star _ _ hf => exact star _ hf + +open Topology in +@[elab_as_elim] +theorem ContinuousMap.induction_on_of_compact {𝕜 : Type*} [RCLike 𝕜] {s : Set 𝕜} [CompactSpace s] + {p : C(s, 𝕜) → Prop} (const : ∀ r, p (.const s r)) (id : p (.restrict s <| .id 𝕜)) + (add : ∀ f g, p f → p g → p (f + g)) (mul : ∀ f g, p f → p g → p (f * g)) + (star : ∀ f, p f → p (star f)) (frequently : ∀ f, (∃ᶠ g in 𝓝 f, p g) → p f) (f : C(s, 𝕜)) : + p f := by + refine f.induction_on const id add mul star fun h f ↦ frequently f ?_ + have := polynomialFunctions.starClosure_topologicalClosure s ▸ mem_top (x := f) + rw [← SetLike.mem_coe, topologicalClosure_coe, mem_closure_iff_frequently] at this + exact this.mp <| .of_forall h + /-- Continuous algebra homomorphisms from `C(s, ℝ)` into an `ℝ`-algebra `A` which agree at `X : 𝕜[X]` (interpreted as a continuous map) are, in fact, equal. -/ @[ext (iff := false)] @@ -548,7 +579,8 @@ lemma ker_evalStarAlgHom_eq_closure_adjoin_id (s : Set 𝕜) (h0 : 0 ∈ s) [Com end ContinuousMap -open ContinuousMapZero in +open scoped ContinuousMapZero + /-- If `s : Set 𝕜` with `RCLike 𝕜` is compact and contains `0`, then the non-unital star subalgebra generated by the identity function in `C(s, 𝕜)₀` is dense. This can be seen as a version of the Weierstrass approximation theorem. -/ @@ -566,4 +598,35 @@ lemma ContinuousMapZero.adjoin_id_dense {s : Set 𝕜} [Zero s] (h0 : ((0 : s) : ContinuousMap.evalStarAlgHom_apply, ContinuousMap.coe_coe] rw [show ⟨0, h0'⟩ = (0 : s) by ext; exact h0.symm, _root_.map_zero f] +/-- An induction principle for `C(s, 𝕜)₀`. -/ +@[elab_as_elim] +lemma ContinuousMapZero.induction_on {s : Set 𝕜} [Zero s] (h0 : ((0 : s) : 𝕜) = 0) + {p : C(s, 𝕜)₀ → Prop} (zero : p 0) (id : p (.id h0)) + (add : ∀ f g, p f → p g → p (f + g)) (mul : ∀ f g, p f → p g → p (f * g)) + (smul : ∀ (r : 𝕜) f, p f → p (r • f)) (star : ∀ f, p f → p (star f)) + (closure : (∀ f ∈ adjoin 𝕜 {(.id h0 : C(s, 𝕜)₀)}, p f) → ∀ f, p f) (f : C(s, 𝕜)₀) : + p f := by + refine closure (fun f hf => ?_) f + induction hf using NonUnitalStarAlgebra.adjoin_induction with + | mem f hf => + simp only [Set.mem_singleton_iff] at hf + rwa [hf] + | zero => exact zero + | add _ _ _ _ hf hg => exact add _ _ hf hg + | mul _ _ _ _ hf hg => exact mul _ _ hf hg + | smul _ _ _ hf => exact smul _ _ hf + | star _ _ hf => exact star _ hf + +open Topology in +@[elab_as_elim] +theorem ContinuousMapZero.induction_on_of_compact {s : Set 𝕜} [Zero s] (h0 : ((0 : s) : 𝕜) = 0) + [CompactSpace s] {p : C(s, 𝕜)₀ → Prop} (zero : p 0) (id : p (.id h0)) + (add : ∀ f g, p f → p g → p (f + g)) (mul : ∀ f g, p f → p g → p (f * g)) + (smul : ∀ (r : 𝕜) f, p f → p (r • f)) (star : ∀ f, p f → p (star f)) + (frequently : ∀ f, (∃ᶠ g in 𝓝 f, p g) → p f) (f : C(s, 𝕜)₀) : + p f := by + refine f.induction_on h0 zero id add mul smul star fun h f ↦ frequently f ?_ + have := (ContinuousMapZero.adjoin_id_dense h0).closure_eq ▸ Set.mem_univ (x := f) + exact mem_closure_iff_frequently.mp this |>.mp <| .of_forall h + end ContinuousMapZero From e7b15d2cf73d42c900766fa0e24e0762f13a8106 Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Sun, 20 Oct 2024 13:19:35 +0000 Subject: [PATCH 203/505] feat(Analysis/Complex/Positivity): new file (#17862) This adds results of the kind "if `f` is an entire function with all iterated derivatives at a point `c` nonnegative real, then `f z` is nonnegative real for `z = c + r` with `r` nonnegative real". From [EulerProducts](https://github.com/MichaelStollBayreuth/EulerProducts). This will be needed for the proof that $L(\chi, 1) \ne 0$ for nontrivial quadratic Dirichlet characters $\chi$ (which in turn is necessary for Dirichlet's Theorem on primes in arithmetic progressions; see [PNT+](https://github.com/AlexKontorovich/PrimeNumberTheoremAnd)). Co-authored-by: Michael Stoll <99838730+MichaelStollBayreuth@users.noreply.github.com> --- Mathlib.lean | 1 + .../Calculus/IteratedDeriv/Lemmas.lean | 42 ++++++++++ Mathlib/Analysis/Complex/Positivity.lean | 81 +++++++++++++++++++ 3 files changed, 124 insertions(+) create mode 100644 Mathlib/Analysis/Complex/Positivity.lean diff --git a/Mathlib.lean b/Mathlib.lean index 0eb264c536463..de1b28985d7b6 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1099,6 +1099,7 @@ import Mathlib.Analysis.Complex.OperatorNorm import Mathlib.Analysis.Complex.PhragmenLindelof import Mathlib.Analysis.Complex.Polynomial.Basic import Mathlib.Analysis.Complex.Polynomial.UnitTrinomial +import Mathlib.Analysis.Complex.Positivity import Mathlib.Analysis.Complex.ReImTopology import Mathlib.Analysis.Complex.RealDeriv import Mathlib.Analysis.Complex.RemovableSingularity diff --git a/Mathlib/Analysis/Calculus/IteratedDeriv/Lemmas.lean b/Mathlib/Analysis/Calculus/IteratedDeriv/Lemmas.lean index 9b8588e207bce..0a4e4120acb15 100644 --- a/Mathlib/Analysis/Calculus/IteratedDeriv/Lemmas.lean +++ b/Mathlib/Analysis/Calculus/IteratedDeriv/Lemmas.lean @@ -15,6 +15,8 @@ This file contains a number of further results on `iteratedDerivWithin` that nee than are available in `Mathlib/Analysis/Calculus/IteratedDeriv/Defs.lean`. -/ +section one_dimensional + variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {F : Type*} [NormedAddCommGroup F] [NormedSpace 𝕜 F] @@ -123,3 +125,43 @@ lemma iteratedDeriv_comp_neg (n : ℕ) (f : 𝕜 → F) (a : 𝕜) : rw [iteratedDeriv_succ, iteratedDeriv_succ, ih', pow_succ', neg_mul, one_mul, deriv_comp_neg (f := fun x ↦ (-1 : 𝕜) ^ n • iteratedDeriv n f x), deriv_const_smul', neg_smul] + +open Topology in +lemma Filter.EventuallyEq.iteratedDeriv_eq (n : ℕ) {f g : 𝕜 → F} {x : 𝕜} (hfg : f =ᶠ[𝓝 x] g) : + iteratedDeriv n f x = iteratedDeriv n g x := by + simp only [← iteratedDerivWithin_univ, iteratedDerivWithin_eq_iteratedFDerivWithin] + rw [(hfg.filter_mono nhdsWithin_le_nhds).iteratedFDerivWithin_eq hfg.eq_of_nhds n] + +lemma Set.EqOn.iteratedDeriv_of_isOpen (hfg : Set.EqOn f g s) (hs : IsOpen s) (n : ℕ) : + Set.EqOn (iteratedDeriv n f) (iteratedDeriv n g) s := by + refine fun x hx ↦ Filter.EventuallyEq.iteratedDeriv_eq n ?_ + filter_upwards [IsOpen.mem_nhds hs hx] with a ha + exact hfg ha + +end one_dimensional + +/-! +### Invariance of iterated derivatives under translation +-/ + +section shift_invariance + +variable {𝕜 F} [NontriviallyNormedField 𝕜] [NormedAddCommGroup F] [NormedSpace 𝕜 F] + +/-- The iterated derivative commutes with shifting the function by a constant on the left. -/ +lemma iteratedDeriv_comp_const_add (n : ℕ) (f : 𝕜 → F) (s : 𝕜) : + iteratedDeriv n (fun z ↦ f (s + z)) = fun t ↦ iteratedDeriv n f (s + t) := by + induction n with + | zero => simp only [iteratedDeriv_zero] + | succ n IH => + simpa only [iteratedDeriv_succ, IH] using funext <| deriv_comp_const_add _ s + +/-- The iterated derivative commutes with shifting the function by a constant on the right. -/ +lemma iteratedDeriv_comp_add_const (n : ℕ) (f : 𝕜 → F) (s : 𝕜) : + iteratedDeriv n (fun z ↦ f (z + s)) = fun t ↦ iteratedDeriv n f (t + s) := by + induction n with + | zero => simp only [iteratedDeriv_zero] + | succ n IH => + simpa only [iteratedDeriv_succ, IH] using funext <| deriv_comp_add_const _ s + +end shift_invariance diff --git a/Mathlib/Analysis/Complex/Positivity.lean b/Mathlib/Analysis/Complex/Positivity.lean new file mode 100644 index 0000000000000..140e3a84386ff --- /dev/null +++ b/Mathlib/Analysis/Complex/Positivity.lean @@ -0,0 +1,81 @@ +/- +Copyright (c) 2024 Michael Stoll. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Michael Stoll +-/ +import Mathlib.Analysis.Complex.TaylorSeries + +/-! +# Nonnegativity of values of holomorphic functions + +We show that if `f` is holomorphic on an open disk `B(c,r)` and all iterated derivatives of `f` +at `c` are nonnegative real, then `f z ≥ 0` for all `z ≥ c` in the disk; see +`DifferentiableOn.nonneg_of_iteratedDeriv_nonneg`. We also provide a +variant `Differentiable.nonneg_of_iteratedDeriv_nonneg` for entire functions and versions +showing `f z ≥ f c` when all iterated derivatives except `f` itseld are nonnegative. +-/ + +open Complex + +open scoped ComplexOrder + +namespace DifferentiableOn + +/-- A function that is holomorphic on the open disk around `c` with radius `r` and whose iterated +derivatives at `c` are all nonnegative real has nonnegative real values on `c + [0,r)`. -/ +theorem nonneg_of_iteratedDeriv_nonneg {f : ℂ → ℂ} {c : ℂ} {r : ℝ} + (hf : DifferentiableOn ℂ f (Metric.ball c r)) (h : ∀ n, 0 ≤ iteratedDeriv n f c) ⦃z : ℂ⦄ + (hz₁ : c ≤ z) (hz₂ : z ∈ Metric.ball c r): + 0 ≤ f z := by + have H := taylorSeries_eq_on_ball' hz₂ hf + rw [← sub_nonneg] at hz₁ + have hz' := eq_re_of_ofReal_le hz₁ + rw [hz'] at hz₁ H + refine H ▸ tsum_nonneg fun n ↦ ?_ + rw [← ofReal_natCast, ← ofReal_pow, ← ofReal_inv, eq_re_of_ofReal_le (h n), ← ofReal_mul, + ← ofReal_mul] + norm_cast at hz₁ ⊢ + have := zero_re ▸ (Complex.le_def.mp (h n)).1 + positivity + +end DifferentiableOn + +namespace Differentiable + +/-- An entire function whose iterated derivatives at `c` are all nonnegative real has nonnegative +real values on `c + ℝ≥0`. -/ +theorem nonneg_of_iteratedDeriv_nonneg {f : ℂ → ℂ} (hf : Differentiable ℂ f) {c : ℂ} + (h : ∀ n, 0 ≤ iteratedDeriv n f c) ⦃z : ℂ⦄ (hz : c ≤ z) : + 0 ≤ f z := by + refine hf.differentiableOn.nonneg_of_iteratedDeriv_nonneg (r := (z - c).re + 1) h hz ?_ + rw [← sub_nonneg] at hz + rw [Metric.mem_ball, dist_eq, eq_re_of_ofReal_le hz] + simpa only [Complex.abs_of_nonneg (nonneg_iff.mp hz).1] using lt_add_one _ + +/-- An entire function whose iterated derivatives at `c` are all nonnegative real (except +possibly the value itself) has values of the form `f c + nonneg. real` on the set `c + ℝ≥0`. -/ +theorem apply_le_of_iteratedDeriv_nonneg {f : ℂ → ℂ} {c : ℂ} (hf : Differentiable ℂ f) + (h : ∀ n ≠ 0, 0 ≤ iteratedDeriv n f c) ⦃z : ℂ⦄ (hz : c ≤ z) : + f c ≤ f z := by + have h' (n : ℕ) : 0 ≤ iteratedDeriv n (f · - f c) c := by + cases n with + | zero => simp only [iteratedDeriv_zero, sub_self, le_refl] + | succ n => + specialize h (n + 1) n.succ_ne_zero + rw [iteratedDeriv_succ'] at h ⊢ + rwa [funext fun x ↦ deriv_sub_const (f := f) (x := x) (f c)] + exact sub_nonneg.mp <| nonneg_of_iteratedDeriv_nonneg (hf.sub_const _) h' hz + +/-- An entire function whose iterated derivatives at `c` are all real with alternating signs +(except possibly the value itself) has values of the form `f c + nonneg. real` along the +set `c - ℝ≥0`. -/ +theorem apply_le_of_iteratedDeriv_alternating {f : ℂ → ℂ} {c : ℂ} (hf : Differentiable ℂ f) + (h : ∀ n ≠ 0, 0 ≤ (-1) ^ n * iteratedDeriv n f c) ⦃z : ℂ⦄ (hz : z ≤ c) : + f c ≤ f z := by + convert apply_le_of_iteratedDeriv_nonneg (f := fun z ↦ f (-z)) + (hf.comp <| differentiable_neg) (fun n hn ↦ ?_) (neg_le_neg_iff.mpr hz) using 1 + · simp only [neg_neg] + · simp only [neg_neg] + · simpa only [iteratedDeriv_comp_neg, neg_neg, smul_eq_mul] using h n hn + +end Differentiable From 0dedd92ed7e30cdcb2c8ff536bed18ea97b79a4f Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 20 Oct 2024 13:29:54 +0000 Subject: [PATCH 204/505] chore: Rename `ClosedEmbedding` to `IsClosedEmbedding` (#17937) `Function.Embedding` is a type while `Embedding` is a proposition, and there are many other kinds of embeddings than topological embeddings. Hence this PR is a step towards 1. renaming `Embedding` to `IsEmbedding` and similarly for neighborhing declarations (which `ClosedEmbedding` is) 2. namespacing it inside `Topology` [Zulip](https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/rename.20.60Inducing.60.20and.20.60Embedding.60.3F). See #15993 for context. --- .../Morphisms/ClosedImmersion.lean | 42 ++++----- .../Morphisms/UnderlyingMap.lean | 14 +-- .../PrimeSpectrum/Basic.lean | 12 ++- Mathlib/AlgebraicGeometry/ResidueField.lean | 2 +- .../Instances.lean | 27 +++--- .../NonUnital.lean | 36 ++++---- .../ContinuousFunctionalCalculus/Order.lean | 4 +- .../Restrict.lean | 46 +++++----- .../ContinuousFunctionalCalculus/Unique.lean | 4 +- .../ContinuousFunctionalCalculus/Unital.lean | 19 ++-- .../Analysis/CStarAlgebra/GelfandDuality.lean | 2 +- .../LineDeriv/IntegrationByParts.lean | 2 +- Mathlib/Analysis/Convex/Intrinsic.lean | 4 +- Mathlib/Analysis/Convolution.lean | 6 +- .../FunctionalSpaces/SobolevInequality.lean | 2 +- Mathlib/Analysis/Normed/Algebra/Spectrum.lean | 4 +- Mathlib/Analysis/Normed/Module/Basic.lean | 4 +- .../Normed/Module/FiniteDimension.lean | 4 +- Mathlib/Analysis/Normed/Operator/Compact.lean | 2 +- .../ContinuousFunctionalCalculus/ExpLog.lean | 2 +- .../SpecialFunctions/Log/ENNRealLogExp.lean | 2 +- Mathlib/CategoryTheory/Galois/Topology.lean | 13 +-- .../Manifold/SmoothManifoldWithCorners.lean | 15 ++-- .../Geometry/Manifold/WhitneyEmbedding.lean | 4 +- .../Matrix/HermitianFunctionalCalculus.lean | 11 ++- .../Constructions/BorelSpace/Basic.lean | 10 ++- .../BorelSpace/ContinuousLinearMap.lean | 4 +- .../Constructions/Polish/EmbeddingReal.lean | 2 +- .../Function/StronglyMeasurable/Basic.lean | 4 +- .../Function/StronglyMeasurable/Lemmas.lean | 2 +- Mathlib/MeasureTheory/Group/Measure.lean | 2 +- Mathlib/MeasureTheory/Integral/Bochner.lean | 7 +- .../Integral/IntervalIntegral.lean | 8 +- .../MeasureTheory/Integral/SetIntegral.lean | 12 ++- .../Measure/Lebesgue/Integral.lean | 2 +- Mathlib/NumberTheory/Modular.lean | 10 +-- .../Algebra/Constructions/DomMulAct.lean | 22 +++-- .../Topology/Algebra/ContinuousMonoidHom.lean | 7 +- .../Algebra/Module/Alternating/Topology.lean | 7 +- .../Algebra/Module/FiniteDimension.lean | 18 ++-- .../Topology/Algebra/ProperAction/Basic.lean | 16 ++-- Mathlib/Topology/Algebra/StarSubalgebra.lean | 18 ++-- .../Category/LightProfinite/Sequence.lean | 9 +- .../Topology/Category/Profinite/Nobeling.lean | 6 +- Mathlib/Topology/Category/Stonean/Limits.lean | 2 +- Mathlib/Topology/Clopen.lean | 2 +- Mathlib/Topology/Compactness/Compact.lean | 28 ++++-- Mathlib/Topology/Compactness/Lindelof.lean | 27 ++++-- .../Topology/Compactness/LocallyCompact.lean | 16 ++-- Mathlib/Topology/Compactness/Paracompact.lean | 9 +- .../Topology/Compactness/SigmaCompact.lean | 9 +- Mathlib/Topology/Constructions.lean | 45 +++++++--- .../ContinuousMap/ContinuousMapZero.lean | 9 +- .../ContinuousMap/StoneWeierstrass.lean | 4 +- Mathlib/Topology/Defs/Induced.lean | 7 +- Mathlib/Topology/DiscreteSubset.lean | 2 +- Mathlib/Topology/FiberBundle/Basic.lean | 7 +- Mathlib/Topology/Homeomorph.lean | 18 ++-- Mathlib/Topology/Instances/CantorSet.lean | 4 +- Mathlib/Topology/Instances/EReal.lean | 5 +- Mathlib/Topology/Instances/Int.lean | 7 +- Mathlib/Topology/Instances/Irrational.lean | 2 +- Mathlib/Topology/Instances/NNReal.lean | 2 +- Mathlib/Topology/Instances/Nat.lean | 7 +- Mathlib/Topology/Instances/Rat.lean | 16 ++-- Mathlib/Topology/Instances/Real.lean | 2 +- Mathlib/Topology/KrullDimension.lean | 13 +-- Mathlib/Topology/LocalAtTarget.lean | 24 ++++-- Mathlib/Topology/Maps/Basic.lean | 54 +++++++----- Mathlib/Topology/Maps/Proper/Basic.lean | 20 +++-- .../Topology/MetricSpace/Antilipschitz.lean | 7 +- Mathlib/Topology/MetricSpace/Basic.lean | 9 +- Mathlib/Topology/MetricSpace/Dilation.lean | 8 +- Mathlib/Topology/MetricSpace/Isometry.lean | 9 +- Mathlib/Topology/MetricSpace/Polish.lean | 19 ++-- Mathlib/Topology/SeparatedMap.lean | 11 ++- Mathlib/Topology/Separation.lean | 53 +++++++----- Mathlib/Topology/Sober.lean | 5 +- Mathlib/Topology/Support.lean | 7 +- Mathlib/Topology/TietzeExtension.lean | 86 ++++++++++++------- Mathlib/Topology/UniformSpace/Ascoli.lean | 21 +++-- .../UniformSpace/CompleteSeparated.lean | 12 ++- .../UniformSpace/UniformEmbedding.lean | 7 +- 83 files changed, 655 insertions(+), 389 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean b/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean index 0e01ff0984171..741e9d4b1146b 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean @@ -40,16 +40,19 @@ namespace AlgebraicGeometry topological map is a closed embedding and the induced stalk maps are surjective. -/ @[mk_iff] class IsClosedImmersion {X Y : Scheme} (f : X ⟶ Y) : Prop where - base_closed : ClosedEmbedding f.base + base_closed : IsClosedEmbedding f.base surj_on_stalks : ∀ x, Function.Surjective (f.stalkMap x) namespace IsClosedImmersion -lemma closedEmbedding {X Y : Scheme} (f : X ⟶ Y) - [IsClosedImmersion f] : ClosedEmbedding f.base := +lemma isClosedEmbedding {X Y : Scheme} (f : X ⟶ Y) + [IsClosedImmersion f] : IsClosedEmbedding f.base := IsClosedImmersion.base_closed -lemma eq_inf : @IsClosedImmersion = (topologically ClosedEmbedding) ⊓ +@[deprecated (since := "2024-10-20")] +alias closedEmbedding := isClosedEmbedding + +lemma eq_inf : @IsClosedImmersion = (topologically IsClosedEmbedding) ⊓ stalkwise (fun f ↦ Function.Surjective f) := by ext X Y f rw [isClosedImmersion_iff] @@ -57,7 +60,7 @@ lemma eq_inf : @IsClosedImmersion = (topologically ClosedEmbedding) ⊓ lemma iff_isPreimmersion {X Y : Scheme} {f : X ⟶ Y} : IsClosedImmersion f ↔ IsPreimmersion f ∧ IsClosed (Set.range f.base) := by - rw [and_comm, isClosedImmersion_iff, isPreimmersion_iff, ← and_assoc, closedEmbedding_iff, + rw [and_comm, isClosedImmersion_iff, isPreimmersion_iff, ← and_assoc, isClosedEmbedding_iff, @and_comm (Embedding _)] lemma of_isPreimmersion {X Y : Scheme} (f : X ⟶ Y) [IsPreimmersion f] @@ -69,7 +72,7 @@ instance (priority := 900) {X Y : Scheme} (f : X ⟶ Y) [IsClosedImmersion f] : /-- Isomorphisms are closed immersions. -/ instance {X Y : Scheme} (f : X ⟶ Y) [IsIso f] : IsClosedImmersion f where - base_closed := Homeomorph.closedEmbedding <| TopCat.homeoOfIso (asIso f.base) + base_closed := Homeomorph.isClosedEmbedding <| TopCat.homeoOfIso (asIso f.base) surj_on_stalks := fun _ ↦ (ConcreteCategory.bijective_of_isIso _).2 instance : MorphismProperty.IsMultiplicative @IsClosedImmersion where @@ -93,7 +96,7 @@ instance respectsIso : MorphismProperty.RespectsIso @IsClosedImmersion := by closed immersion. -/ theorem spec_of_surjective {R S : CommRingCat} (f : R ⟶ S) (h : Function.Surjective f) : IsClosedImmersion (Spec.map f) where - base_closed := PrimeSpectrum.closedEmbedding_comap_of_surjective _ _ h + base_closed := PrimeSpectrum.isClosedEmbedding_comap_of_surjective _ _ h surj_on_stalks x := by haveI : (RingHom.toMorphismProperty (fun f ↦ Function.Surjective f)).RespectsIso := by rw [← RingHom.toMorphismProperty_respectsIso_iff] @@ -121,14 +124,13 @@ lemma of_surjective_of_isAffine {X Y : Scheme} [IsAffine X] [IsAffine Y] (f : X theorem of_comp {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) [IsClosedImmersion g] [IsClosedImmersion (f ≫ g)] : IsClosedImmersion f where base_closed := by - have h := closedEmbedding (f ≫ g) - rw [Scheme.comp_base] at h - apply closedEmbedding_of_continuous_injective_closed (Scheme.Hom.continuous f) - · exact Function.Injective.of_comp h.inj - · intro Z hZ - rw [ClosedEmbedding.closed_iff_image_closed (closedEmbedding g), - ← Set.image_comp] - exact ClosedEmbedding.isClosedMap h _ hZ + have h := isClosedEmbedding (f ≫ g) + simp only [Scheme.comp_coeBase, TopCat.coe_comp] at h + refine .of_continuous_injective_isClosedMap (Scheme.Hom.continuous f) h.inj.of_comp ?_ + intro Z hZ + rw [IsClosedEmbedding.closed_iff_image_closed (isClosedEmbedding g), + ← Set.image_comp] + exact h.isClosedMap _ hZ surj_on_stalks x := by have h := (f ≫ g).stalkMap_surjective x simp_rw [Scheme.stalkMap_comp] at h @@ -222,14 +224,14 @@ namespace IsClosedImmersion sections is injective, `f` is an isomorphism. -/ theorem isIso_of_injective_of_isAffine [IsClosedImmersion f] (hf : Function.Injective (f.app ⊤)) : IsIso f := (isIso_iff_stalk_iso f).mpr <| - have : CompactSpace X := (closedEmbedding f).compactSpace + have : CompactSpace X := (isClosedEmbedding f).compactSpace have hiso : IsIso f.base := TopCat.isIso_of_bijective_of_isClosedMap _ - ⟨(closedEmbedding f).inj, - surjective_of_isClosed_range_of_injective ((closedEmbedding f).isClosed_range) hf⟩ - ((closedEmbedding f).isClosedMap) + ⟨(isClosedEmbedding f).inj, + surjective_of_isClosed_range_of_injective ((isClosedEmbedding f).isClosed_range) hf⟩ + ((isClosedEmbedding f).isClosedMap) ⟨hiso, fun x ↦ (ConcreteCategory.isIso_iff_bijective _).mpr ⟨stalkMap_injective_of_isOpenMap_of_injective ((TopCat.homeoOfIso (asIso f.base)).isOpenMap) - (closedEmbedding f).inj hf _, f.stalkMap_surjective x⟩⟩ + (isClosedEmbedding f).inj hf _, f.stalkMap_surjective x⟩⟩ variable (f) diff --git a/Mathlib/AlgebraicGeometry/Morphisms/UnderlyingMap.lean b/Mathlib/AlgebraicGeometry/Morphisms/UnderlyingMap.lean index 00d8b44dc43f0..ed2f3268b2efc 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/UnderlyingMap.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/UnderlyingMap.lean @@ -19,7 +19,7 @@ of the underlying map of topological spaces, including - `IsClosedMap` - `Embedding` - `IsOpenEmbedding` -- `ClosedEmbedding` +- `IsClosedEmbedding` -/ @@ -131,16 +131,16 @@ instance isOpenEmbedding_isLocalAtTarget : IsLocalAtTarget (topologically IsOpen end IsOpenEmbedding -section ClosedEmbedding +section IsClosedEmbedding -instance : (topologically ClosedEmbedding).RespectsIso := - topologically_respectsIso _ (fun e ↦ e.closedEmbedding) (fun _ _ hf hg ↦ hg.comp hf) +instance : (topologically IsClosedEmbedding).RespectsIso := + topologically_respectsIso _ (fun e ↦ e.isClosedEmbedding) (fun _ _ hf hg ↦ hg.comp hf) -instance closedEmbedding_isLocalAtTarget : IsLocalAtTarget (topologically ClosedEmbedding) := +instance isClosedEmbedding_isLocalAtTarget : IsLocalAtTarget (topologically IsClosedEmbedding) := topologically_isLocalAtTarget _ (fun _ s hf ↦ hf.restrictPreimage s) - (fun _ _ _ hU hfcont hf ↦ (closedEmbedding_iff_closedEmbedding_of_iSup_eq_top hU hfcont).mpr hf) + (fun _ _ _ hU hfcont ↦ (isClosedEmbedding_iff_isClosedEmbedding_of_iSup_eq_top hU hfcont).mpr) -end ClosedEmbedding +end IsClosedEmbedding end AlgebraicGeometry diff --git a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean index f918a39487b28..2ad92e90d3efa 100644 --- a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean +++ b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean @@ -326,11 +326,14 @@ theorem isClosed_range_comap_of_surjective (hf : Surjective f) : rw [range_comap_of_surjective _ f hf] exact isClosed_zeroLocus _ -theorem closedEmbedding_comap_of_surjective (hf : Surjective f) : ClosedEmbedding (comap f) := +theorem isClosedEmbedding_comap_of_surjective (hf : Surjective f) : IsClosedEmbedding (comap f) := { induced := (comap_inducing_of_surjective S f hf).induced inj := comap_injective_of_surjective f hf isClosed_range := isClosed_range_comap_of_surjective S f hf } +@[deprecated (since := "2024-10-20")] +alias closedEmbedding_comap_of_surjective := isClosedEmbedding_comap_of_surjective + end SpecOfSurjective section SpecProd @@ -351,17 +354,18 @@ noncomputable def primeSpectrumProdHomeo : PrimeSpectrum (R × S) ≃ₜ PrimeSpectrum R ⊕ PrimeSpectrum S := by refine ((primeSpectrumProd R S).symm.toHomeomorphOfInducing ?_).symm - refine (closedEmbedding_of_continuous_injective_closed ?_ (Equiv.injective _) ?_).toInducing + refine (IsClosedEmbedding.of_continuous_injective_isClosedMap ?_ + (Equiv.injective _) ?_).toInducing · rw [continuous_sum_dom] simp only [Function.comp_def, primeSpectrumProd_symm_inl, primeSpectrumProd_symm_inr] exact ⟨(comap _).2, (comap _).2⟩ · rw [isClosedMap_sum] constructor · simp_rw [primeSpectrumProd_symm_inl] - refine (closedEmbedding_comap_of_surjective _ _ ?_).isClosedMap + refine (isClosedEmbedding_comap_of_surjective _ _ ?_).isClosedMap exact Prod.fst_surjective · simp_rw [primeSpectrumProd_symm_inr] - refine (closedEmbedding_comap_of_surjective _ _ ?_).isClosedMap + refine (isClosedEmbedding_comap_of_surjective _ _ ?_).isClosedMap exact Prod.snd_surjective end SpecProd diff --git a/Mathlib/AlgebraicGeometry/ResidueField.lean b/Mathlib/AlgebraicGeometry/ResidueField.lean index c7fcb54274fd0..fd0360bc3bc0c 100644 --- a/Mathlib/AlgebraicGeometry/ResidueField.lean +++ b/Mathlib/AlgebraicGeometry/ResidueField.lean @@ -52,7 +52,7 @@ def residue (X : Scheme.{u}) (x) : X.presheaf.stalk x ⟶ X.residueField x := `Spec.map (X.residue x)` is a closed immersion. -/ instance {X : Scheme.{u}} (x) : IsPreimmersion (Spec.map (X.residue x)) := IsPreimmersion.mk_Spec_map - (PrimeSpectrum.closedEmbedding_comap_of_surjective _ _ (Ideal.Quotient.mk_surjective)).embedding + (PrimeSpectrum.isClosedEmbedding_comap_of_surjective _ _ Ideal.Quotient.mk_surjective).embedding (RingHom.surjectiveOnStalks_of_surjective (Ideal.Quotient.mk_surjective)) @[simp] diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean index ea67a1f84c970..5d998afdff365 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean @@ -78,16 +78,19 @@ noncomputable def cfcₙAux : C(σₙ 𝕜 a, 𝕜)₀ →⋆ₙₐ[𝕜] A⁺¹ lemma cfcₙAux_id : cfcₙAux hp₁ a ha (ContinuousMapZero.id rfl) = a := cfcHom_id (hp₁.mpr ha) open Unitization in -lemma closedEmbedding_cfcₙAux : ClosedEmbedding (cfcₙAux hp₁ a ha) := by +lemma isClosedEmbedding_cfcₙAux : IsClosedEmbedding (cfcₙAux hp₁ a ha) := by simp only [cfcₙAux, NonUnitalStarAlgHom.coe_comp] - refine ((cfcHom_closedEmbedding (hp₁.mpr ha)).comp ?_).comp - ContinuousMapZero.closedEmbedding_toContinuousMap + refine ((cfcHom_isClosedEmbedding (hp₁.mpr ha)).comp ?_).comp + ContinuousMapZero.isClosedEmbedding_toContinuousMap let e : C(σₙ 𝕜 a, 𝕜) ≃ₜ C(σ 𝕜 (a : A⁺¹), 𝕜) := { (Homeomorph.compStarAlgEquiv' 𝕜 𝕜 <| .setCongr <| (quasispectrum_eq_spectrum_inr' 𝕜 𝕜 a).symm) with continuous_toFun := ContinuousMap.continuous_comp_left _ continuous_invFun := ContinuousMap.continuous_comp_left _ } - exact e.closedEmbedding + exact e.isClosedEmbedding + +@[deprecated (since := "2024-10-20")] +alias closedEmbedding_cfcₙAux := isClosedEmbedding_cfcₙAux lemma spec_cfcₙAux (f : C(σₙ 𝕜 a, 𝕜)₀) : σ 𝕜 (cfcₙAux hp₁ a ha f) = Set.range f := by rw [cfcₙAux, NonUnitalStarAlgHom.comp_assoc, NonUnitalStarAlgHom.comp_apply] @@ -103,7 +106,7 @@ variable [CompleteSpace A] lemma cfcₙAux_mem_range_inr (f : C(σₙ 𝕜 a, 𝕜)₀) : cfcₙAux hp₁ a ha f ∈ NonUnitalStarAlgHom.range (Unitization.inrNonUnitalStarAlgHom 𝕜 A) := by - have h₁ := (closedEmbedding_cfcₙAux hp₁ a ha).continuous.range_subset_closure_image_dense + have h₁ := (isClosedEmbedding_cfcₙAux hp₁ a ha).continuous.range_subset_closure_image_dense (ContinuousMapZero.adjoin_id_dense (s := σₙ 𝕜 a) rfl) ⟨f, rfl⟩ rw [← SetLike.mem_coe] refine closure_minimal ?_ ?_ h₁ @@ -133,11 +136,11 @@ theorem RCLike.nonUnitalContinuousFunctionalCalculus : have coe_ψ (f : C(σₙ 𝕜 a, 𝕜)₀) : ψ f = cfcₙAux hp₁ a ha f := congr_arg Subtype.val <| (inrRangeEquiv 𝕜 A).apply_symm_apply ⟨cfcₙAux hp₁ a ha f, cfcₙAux_mem_range_inr hp₁ a ha f⟩ - refine ⟨ψ, ?closedEmbedding, ?map_id, fun f ↦ ?map_spec, fun f ↦ ?isStarNormal⟩ - case closedEmbedding => - apply isometry_inr (𝕜 := 𝕜) (A := A) |>.closedEmbedding |>.of_comp_iff.mp + refine ⟨ψ, ?isClosedEmbedding, ?map_id, fun f ↦ ?map_spec, fun f ↦ ?isStarNormal⟩ + case isClosedEmbedding => + apply isometry_inr (𝕜 := 𝕜) (A := A) |>.isClosedEmbedding |>.of_comp_iff.mp have : inr ∘ ψ = cfcₙAux hp₁ a ha := by ext1; rw [Function.comp_apply, coe_ψ] - exact this ▸ closedEmbedding_cfcₙAux hp₁ a ha + exact this ▸ isClosedEmbedding_cfcₙAux hp₁ a ha case map_id => exact inr_injective (R := 𝕜) <| coe_ψ _ ▸ cfcₙAux_id hp₁ a ha case map_spec => exact quasispectrum_eq_spectrum_inr' 𝕜 𝕜 (ψ f) ▸ coe_ψ _ ▸ spec_cfcₙAux hp₁ a ha f @@ -158,13 +161,13 @@ instance IsStarNormal.instContinuousFunctionalCalculus {A : Type*} [NormedRing A spectrum_nonempty a _ := spectrum.nonempty a exists_cfc_of_predicate a ha := by refine ⟨(elementalStarAlgebra ℂ a).subtype.comp <| continuousFunctionalCalculus a, - ?hom_closedEmbedding, ?hom_id, ?hom_map_spectrum, ?predicate_hom⟩ - case hom_closedEmbedding => + ?hom_isClosedEmbedding, ?hom_id, ?hom_map_spectrum, ?predicate_hom⟩ + case hom_isClosedEmbedding => -- note: Lean should find these for `StarAlgEquiv.isometry`, but it doesn't and so we -- provide them manually. have : SMulCommClass ℂ C(σ ℂ a, ℂ) C(σ ℂ a, ℂ) := Algebra.to_smulCommClass (A := C(σ ℂ a, ℂ)) have : IsScalarTower ℂ C(σ ℂ a, ℂ) C(σ ℂ a, ℂ) := IsScalarTower.right (A := C(σ ℂ a, ℂ)) - exact Isometry.closedEmbedding <| + exact Isometry.isClosedEmbedding <| isometry_subtype_coe.comp <| StarAlgEquiv.isometry (continuousFunctionalCalculus a) case hom_id => exact congr_arg Subtype.val <| continuousFunctionalCalculus_map_id a case hom_map_spectrum => diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean index d4a2d892c661e..e837b34df4812 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean @@ -75,7 +75,7 @@ class NonUnitalContinuousFunctionalCalculus (R : Type*) {A : Type*} (p : outPara predicate_zero : p 0 [compactSpace_quasispectrum : ∀ a : A, CompactSpace (σₙ R a)] exists_cfc_of_predicate : ∀ a, p a → ∃ φ : C(σₙ R a, R)₀ →⋆ₙₐ[R] A, - ClosedEmbedding φ ∧ φ ⟨(ContinuousMap.id R).restrict <| σₙ R a, rfl⟩ = a ∧ + IsClosedEmbedding φ ∧ φ ⟨(ContinuousMap.id R).restrict <| σₙ R a, rfl⟩ = a ∧ (∀ f, σₙ R (φ f) = Set.range f) ∧ ∀ f, p (φ f) -- this instance should not be activated everywhere but it is useful when developing generic API @@ -128,13 +128,16 @@ the user should instead prefer `cfcₙ` over `cfcₙHom`. noncomputable def cfcₙHom : C(σₙ R a, R)₀ →⋆ₙₐ[R] A := (NonUnitalContinuousFunctionalCalculus.exists_cfc_of_predicate a ha).choose -lemma cfcₙHom_closedEmbedding : - ClosedEmbedding <| (cfcₙHom ha : C(σₙ R a, R)₀ →⋆ₙₐ[R] A) := +lemma cfcₙHom_isClosedEmbedding : + IsClosedEmbedding <| (cfcₙHom ha : C(σₙ R a, R)₀ →⋆ₙₐ[R] A) := (NonUnitalContinuousFunctionalCalculus.exists_cfc_of_predicate a ha).choose_spec.1 +@[deprecated (since := "2024-10-20")] +alias cfcₙHom_closedEmbedding := cfcₙHom_isClosedEmbedding + @[fun_prop] lemma cfcₙHom_continuous : Continuous (cfcₙHom ha : C(σₙ R a, R)₀ →⋆ₙₐ[R] A) := - cfcₙHom_closedEmbedding ha |>.continuous + cfcₙHom_isClosedEmbedding ha |>.continuous lemma cfcₙHom_id : cfcₙHom ha ⟨(ContinuousMap.id R).restrict <| σₙ R a, rfl⟩ = a := @@ -152,7 +155,7 @@ lemma cfcₙHom_predicate (f : C(σₙ R a, R)₀) : lemma cfcₙHom_eq_of_continuous_of_map_id [UniqueNonUnitalContinuousFunctionalCalculus R A] (φ : C(σₙ R a, R)₀ →⋆ₙₐ[R] A) (hφ₁ : Continuous φ) (hφ₂ : φ ⟨.restrict (σₙ R a) <| .id R, rfl⟩ = a) : cfcₙHom ha = φ := - (cfcₙHom ha).ext_continuousMap a φ (cfcₙHom_closedEmbedding ha).continuous hφ₁ <| by + (cfcₙHom ha).ext_continuousMap a φ (cfcₙHom_isClosedEmbedding ha).continuous hφ₁ <| by rw [cfcₙHom_id ha, hφ₂] theorem cfcₙHom_comp [UniqueNonUnitalContinuousFunctionalCalculus R A] (f : C(σₙ R a, R)₀) @@ -169,7 +172,7 @@ theorem cfcₙHom_comp [UniqueNonUnitalContinuousFunctionalCalculus R A] (f : C( let φ : C(σₙ R (cfcₙHom ha f), R)₀ →⋆ₙₐ[R] A := (cfcₙHom ha).comp ψ suffices cfcₙHom (cfcₙHom_predicate ha f) = φ from DFunLike.congr_fun this.symm g refine cfcₙHom_eq_of_continuous_of_map_id (cfcₙHom_predicate ha f) φ ?_ ?_ - · refine (cfcₙHom_closedEmbedding ha).continuous.comp <| continuous_induced_rng.mpr ?_ + · refine (cfcₙHom_isClosedEmbedding ha).continuous.comp <| continuous_induced_rng.mpr ?_ exact f'.toContinuousMap.continuous_comp_left.comp continuous_induced_dom · simp only [φ, ψ, NonUnitalStarAlgHom.comp_apply, NonUnitalStarAlgHom.coe_mk', NonUnitalAlgHom.coe_mk] @@ -187,7 +190,7 @@ noncomputable def cfcₙL {a : A} (ha : p a) : C(σₙ R a, R)₀ →L[R] A := { cfcₙHom ha with toFun := cfcₙHom ha map_smul' := map_smul _ - cont := (cfcₙHom_closedEmbedding ha).continuous } + cont := (cfcₙHom_isClosedEmbedding ha).continuous } end cfcₙL @@ -306,7 +309,7 @@ lemma eqOn_of_cfcₙ_eq_cfcₙ {f g : R → R} {a : A} (h : cfcₙ f a = cfcₙ (hg : ContinuousOn g (σₙ R a) := by cfc_cont_tac) (hg0 : g 0 = 0 := by cfc_zero_tac) : (σₙ R a).EqOn f g := by rw [cfcₙ_apply f a, cfcₙ_apply g a] at h - have := (cfcₙHom_closedEmbedding (show p a from ha) (R := R)).inj h + have := (cfcₙHom_isClosedEmbedding (show p a from ha) (R := R)).inj h intro x hx congrm($(this) ⟨x, hx⟩) @@ -695,12 +698,12 @@ variable [CompleteSpace R] -- gives access to the `ContinuousFunctionalCalculus.compactSpace_spectrum` instance open scoped ContinuousFunctionalCalculus -lemma closedEmbedding_cfcₙHom_of_cfcHom {a : A} (ha : p a) : - ClosedEmbedding (cfcₙHom_of_cfcHom R ha) := by +lemma isClosedEmbedding_cfcₙHom_of_cfcHom {a : A} (ha : p a) : + IsClosedEmbedding (cfcₙHom_of_cfcHom R ha) := by let f : C(spectrum R a, σₙ R a) := ⟨_, continuous_inclusion <| spectrum_subset_quasispectrum R a⟩ - refine (cfcHom_closedEmbedding ha).comp <| - (IsUniformInducing.isUniformEmbedding ⟨?_⟩).toClosedEmbedding + refine (cfcHom_isClosedEmbedding ha).comp <| + (IsUniformInducing.isUniformEmbedding ⟨?_⟩).toIsClosedEmbedding have := uniformSpace_eq_inf_precomp_of_cover (β := R) f (0 : C(Unit, σₙ R a)) (map_continuous f).isProperMap (map_continuous 0).isProperMap <| by simp only [← Subtype.val_injective.image_injective.eq_iff, f, ContinuousMap.coe_mk, @@ -715,6 +718,9 @@ lemma closedEmbedding_cfcₙHom_of_cfcHom {a : A} (ha : p a) : convert Filter.comap_const_of_mem this with ⟨u, v⟩ <;> ext ⟨x, rfl⟩ <;> [exact map_zero u; exact map_zero v] +@[deprecated (since := "2024-10-20")] +alias closedEmbedding_cfcₙHom_of_cfcHom := isClosedEmbedding_cfcₙHom_of_cfcHom + instance ContinuousFunctionalCalculus.toNonUnital : NonUnitalContinuousFunctionalCalculus R p where predicate_zero := cfc_predicate_zero R compactSpace_quasispectrum a := by @@ -723,7 +729,7 @@ instance ContinuousFunctionalCalculus.toNonUnital : NonUnitalContinuousFunctiona exact h_cpct |>.union isCompact_singleton exists_cfc_of_predicate _ ha := ⟨cfcₙHom_of_cfcHom R ha, - closedEmbedding_cfcₙHom_of_cfcHom ha, + isClosedEmbedding_cfcₙHom_of_cfcHom ha, cfcHom_id ha, cfcₙHom_of_cfcHom_map_quasispectrum ha, fun _ ↦ cfcHom_predicate ha _⟩ @@ -734,8 +740,8 @@ lemma cfcₙHom_eq_cfcₙHom_of_cfcHom [UniqueNonUnitalContinuousFunctionalCalcu refine UniqueNonUnitalContinuousFunctionalCalculus.eq_of_continuous_of_map_id (σₙ R a) ?_ _ _ ?_ ?_ ?_ · simp - · exact (cfcₙHom_closedEmbedding (R := R) ha).continuous - · exact (closedEmbedding_cfcₙHom_of_cfcHom ha).continuous + · exact (cfcₙHom_isClosedEmbedding (R := R) ha).continuous + · exact (isClosedEmbedding_cfcₙHom_of_cfcHom ha).continuous · simpa only [cfcₙHom_id (R := R) ha] using (cfcHom_id ha).symm /-- When `cfc` is applied to a function that maps zero to zero, it is equivalent to using diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Order.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Order.lean index 10e0978f35a5c..08951f013e6d4 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Order.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Order.lean @@ -370,8 +370,8 @@ lemma conjugate_le_norm_smul' {a b : A} (hb : IsSelfAdjoint b := by cfc_tac) : /-- The set of nonnegative elements in a C⋆-algebra is closed. -/ lemma isClosed_nonneg : IsClosed {a : A | 0 ≤ a} := by suffices IsClosed {a : Unitization ℂ A | 0 ≤ a} by - rw [Unitization.isometry_inr (𝕜 := ℂ) |>.closedEmbedding.closed_iff_image_closed] - convert this.inter <| (Unitization.isometry_inr (𝕜 := ℂ)).closedEmbedding.isClosed_range + rw [Unitization.isometry_inr (𝕜 := ℂ) |>.isClosedEmbedding.closed_iff_image_closed] + convert this.inter <| (Unitization.isometry_inr (𝕜 := ℂ)).isClosedEmbedding.isClosed_range ext a simp only [Set.mem_image, Set.mem_setOf_eq, Set.mem_inter_iff, Set.mem_range, ← exists_and_left] congr! 2 with x diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Restrict.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Restrict.lean index bb43228df3fdf..ac2ac678b6bee 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Restrict.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Restrict.lean @@ -82,14 +82,17 @@ lemma starAlgHom_id {a : A} {φ : C(spectrum S a, S) →⋆ₐ[S] A} {f : C(S, R variable [TopologicalSpace A] [ContinuousFunctionalCalculus S q] variable [CompleteSpace R] -lemma closedEmbedding_starAlgHom {a : A} {φ : C(spectrum S a, S) →⋆ₐ[S] A} - (hφ : ClosedEmbedding φ) {f : C(S, R)} (h : SpectrumRestricts a f) +lemma isClosedEmbedding_starAlgHom {a : A} {φ : C(spectrum S a, S) →⋆ₐ[S] A} + (hφ : IsClosedEmbedding φ) {f : C(S, R)} (h : SpectrumRestricts a f) (halg : IsUniformEmbedding (algebraMap R S)) : - ClosedEmbedding (h.starAlgHom φ) := - hφ.comp <| IsUniformEmbedding.toClosedEmbedding <| .comp + IsClosedEmbedding (h.starAlgHom φ) := + hφ.comp <| IsUniformEmbedding.toIsClosedEmbedding <| .comp (ContinuousMap.isUniformEmbedding_comp _ halg) (UniformEquiv.arrowCongr h.homeomorph.symm (.refl _) |>.isUniformEmbedding) +@[deprecated (since := "2024-10-20")] +alias closedEmbedding_starAlgHom := isClosedEmbedding_starAlgHom + /-- Given a `ContinuousFunctionalCalculus S q`. If we form the predicate `p` for `a : A` characterized by: `q a` and the spectrum of `a` restricts to the scalar subring `R` via `f : C(S, R)`, then we can get a restricted functional calculus @@ -103,13 +106,13 @@ protected theorem cfc (f : C(S, R)) (halg : IsUniformEmbedding (algebraMap R S)) compactSpace_spectrum a := by have := ContinuousFunctionalCalculus.compactSpace_spectrum (R := S) a rw [← isCompact_iff_compactSpace] at this ⊢ - simpa using halg.toClosedEmbedding.isCompact_preimage this + simpa using halg.toIsClosedEmbedding.isCompact_preimage this exists_cfc_of_predicate a ha := by refine ⟨((h a).mp ha).2.starAlgHom (cfcHom ((h a).mp ha).1 (R := S)), - ?hom_closedEmbedding, ?hom_id, ?hom_map_spectrum, ?predicate_hom⟩ - case hom_closedEmbedding => - exact ((h a).mp ha).2.closedEmbedding_starAlgHom - (cfcHom_closedEmbedding ((h a).mp ha).1) halg + ?hom_isClosedEmbedding, ?hom_id, ?hom_map_spectrum, ?predicate_hom⟩ + case hom_isClosedEmbedding => + exact ((h a).mp ha).2.isClosedEmbedding_starAlgHom + (cfcHom_isClosedEmbedding ((h a).mp ha).1) halg case hom_id => exact ((h a).mp ha).2.starAlgHom_id <| cfcHom_id ((h a).mp ha).1 case hom_map_spectrum => intro g @@ -143,7 +146,7 @@ lemma cfcHom_eq_restrict (f : C(S, R)) (halg : IsUniformEmbedding (algebraMap R {a : A} (hpa : p a) (hqa : q a) (h : SpectrumRestricts a f) : cfcHom hpa = h.starAlgHom (cfcHom hqa) := by apply cfcHom_eq_of_continuous_of_map_id - · exact h.closedEmbedding_starAlgHom (cfcHom_closedEmbedding hqa) halg |>.continuous + · exact h.isClosedEmbedding_starAlgHom (cfcHom_isClosedEmbedding hqa) halg |>.continuous · exact h.starAlgHom_id (cfcHom_id hqa) lemma cfc_eq_restrict (f : C(S, R)) (halg : IsUniformEmbedding (algebraMap R S)) {a : A} (hpa : p a) @@ -222,15 +225,18 @@ lemma nonUnitalStarAlgHom_id {a : A} {φ : C(σₙ S a, S)₀ →⋆ₙₐ[S] A} variable [TopologicalSpace A] [NonUnitalContinuousFunctionalCalculus S q] variable [CompleteSpace R] -lemma closedEmbedding_nonUnitalStarAlgHom {a : A} {φ : C(σₙ S a, S)₀ →⋆ₙₐ[S] A} - (hφ : ClosedEmbedding φ) {f : C(S, R)} (h : QuasispectrumRestricts a f) +lemma isClosedEmbedding_nonUnitalStarAlgHom {a : A} {φ : C(σₙ S a, S)₀ →⋆ₙₐ[S] A} + (hφ : IsClosedEmbedding φ) {f : C(S, R)} (h : QuasispectrumRestricts a f) (halg : IsUniformEmbedding (algebraMap R S)) : - ClosedEmbedding (h.nonUnitalStarAlgHom φ) := by + IsClosedEmbedding (h.nonUnitalStarAlgHom φ) := by have : h.homeomorph.symm 0 = 0 := Subtype.ext (map_zero <| algebraMap _ _) - refine hφ.comp <| IsUniformEmbedding.toClosedEmbedding <| .comp + refine hφ.comp <| IsUniformEmbedding.toIsClosedEmbedding <| .comp (ContinuousMapZero.isUniformEmbedding_comp _ halg) (UniformEquiv.arrowCongrLeft₀ h.homeomorph.symm this |>.isUniformEmbedding) +@[deprecated (since := "2024-10-20")] +alias closedEmbedding_nonUnitalStarAlgHom := isClosedEmbedding_nonUnitalStarAlgHom + variable [IsScalarTower R A A] [SMulCommClass R A A] /-- Given a `NonUnitalContinuousFunctionalCalculus S q`. If we form the predicate `p` for `a : A` @@ -244,13 +250,13 @@ protected theorem cfc (f : C(S, R)) (halg : IsUniformEmbedding (algebraMap R S)) compactSpace_quasispectrum a := by have := NonUnitalContinuousFunctionalCalculus.compactSpace_quasispectrum (R := S) a rw [← isCompact_iff_compactSpace] at this ⊢ - simpa using halg.toClosedEmbedding.isCompact_preimage this + simpa using halg.toIsClosedEmbedding.isCompact_preimage this exists_cfc_of_predicate a ha := by refine ⟨((h a).mp ha).2.nonUnitalStarAlgHom (cfcₙHom ((h a).mp ha).1 (R := S)), - ?hom_closedEmbedding, ?hom_id, ?hom_map_spectrum, ?predicate_hom⟩ - case hom_closedEmbedding => - exact ((h a).mp ha).2.closedEmbedding_nonUnitalStarAlgHom - (cfcₙHom_closedEmbedding ((h a).mp ha).1) halg + ?hom_isClosedEmbedding, ?hom_id, ?hom_map_spectrum, ?predicate_hom⟩ + case hom_isClosedEmbedding => + exact ((h a).mp ha).2.isClosedEmbedding_nonUnitalStarAlgHom + (cfcₙHom_isClosedEmbedding ((h a).mp ha).1) halg case hom_id => exact ((h a).mp ha).2.nonUnitalStarAlgHom_id <| cfcₙHom_id ((h a).mp ha).1 case hom_map_spectrum => intro g @@ -289,7 +295,7 @@ lemma cfcₙHom_eq_restrict (f : C(S, R)) (halg : IsUniformEmbedding (algebraMap (hpa : p a) (hqa : q a) (h : QuasispectrumRestricts a f) : cfcₙHom hpa = h.nonUnitalStarAlgHom (cfcₙHom hqa) := by apply cfcₙHom_eq_of_continuous_of_map_id - · exact h.closedEmbedding_nonUnitalStarAlgHom (cfcₙHom_closedEmbedding hqa) halg |>.continuous + · exact h.isClosedEmbedding_nonUnitalStarAlgHom (cfcₙHom_isClosedEmbedding hqa) halg |>.continuous · exact h.nonUnitalStarAlgHom_id (cfcₙHom_id hqa) lemma cfcₙ_eq_restrict (f : C(S, R)) (halg : IsUniformEmbedding (algebraMap R S)) {a : A} diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unique.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unique.lean index 00f8c9bf97f18..5186931925074 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unique.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unique.lean @@ -169,7 +169,7 @@ instance NNReal.instUniqueContinuousFunctionalCalculus [UniqueContinuousFunction have : CompactSpace (spectrum ℝ a) := UniqueContinuousFunctionalCalculus.compactSpace_spectrum a rw [← isCompact_iff_compactSpace] at * rw [← spectrum.preimage_algebraMap ℝ] - exact closedEmbedding_subtype_val isClosed_nonneg |>.isCompact_preimage <| by assumption + exact isClosed_nonneg.isClosedEmbedding_subtypeVal.isCompact_preimage <| by assumption eq_of_continuous_of_map_id s hs φ ψ hφ hψ h := by let s' : Set ℝ := (↑) '' s let e : s ≃ₜ s' := @@ -376,7 +376,7 @@ instance NNReal.instUniqueNonUnitalContinuousFunctionalCalculus UniqueNonUnitalContinuousFunctionalCalculus.compactSpace_quasispectrum a rw [← isCompact_iff_compactSpace] at * rw [← quasispectrum.preimage_algebraMap ℝ] - exact closedEmbedding_subtype_val isClosed_nonneg |>.isCompact_preimage <| by assumption + exact isClosed_nonneg.isClosedEmbedding_subtypeVal.isCompact_preimage <| by assumption eq_of_continuous_of_map_id s hs _inst h0 φ ψ hφ hψ h := by let s' : Set ℝ := (↑) '' s let e : s ≃ₜ s' := diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean index 29fa35441c946..c6070f8408685 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean @@ -165,7 +165,7 @@ class ContinuousFunctionalCalculus (R : Type*) {A : Type*} (p : outParam (A → [compactSpace_spectrum (a : A) : CompactSpace (spectrum R a)] spectrum_nonempty [Nontrivial A] (a : A) (ha : p a) : (spectrum R a).Nonempty exists_cfc_of_predicate : ∀ a, p a → ∃ φ : C(spectrum R a, R) →⋆ₐ[R] A, - ClosedEmbedding φ ∧ φ ((ContinuousMap.id R).restrict <| spectrum R a) = a ∧ + IsClosedEmbedding φ ∧ φ ((ContinuousMap.id R).restrict <| spectrum R a) = a ∧ (∀ f, spectrum R (φ f) = Set.range f) ∧ ∀ f, p (φ f) -- this instance should not be activated everywhere but it is useful when developing generic API @@ -228,13 +228,16 @@ user should instead prefer `cfc` over `cfcHom`. noncomputable def cfcHom : C(spectrum R a, R) →⋆ₐ[R] A := (ContinuousFunctionalCalculus.exists_cfc_of_predicate a ha).choose -lemma cfcHom_closedEmbedding : - ClosedEmbedding <| (cfcHom ha : C(spectrum R a, R) →⋆ₐ[R] A) := +lemma cfcHom_isClosedEmbedding : + IsClosedEmbedding <| (cfcHom ha : C(spectrum R a, R) →⋆ₐ[R] A) := (ContinuousFunctionalCalculus.exists_cfc_of_predicate a ha).choose_spec.1 +@[deprecated (since := "2024-10-20")] +alias cfcHom_closedEmbedding := cfcHom_isClosedEmbedding + @[fun_prop] lemma cfcHom_continuous : Continuous (cfcHom ha : C(spectrum R a, R) →⋆ₐ[R] A) := - cfcHom_closedEmbedding ha |>.continuous + cfcHom_isClosedEmbedding ha |>.continuous lemma cfcHom_id : cfcHom ha ((ContinuousMap.id R).restrict <| spectrum R a) = a := @@ -252,7 +255,7 @@ lemma cfcHom_predicate (f : C(spectrum R a, R)) : lemma cfcHom_eq_of_continuous_of_map_id [UniqueContinuousFunctionalCalculus R A] (φ : C(spectrum R a, R) →⋆ₐ[R] A) (hφ₁ : Continuous φ) (hφ₂ : φ (.restrict (spectrum R a) <| .id R) = a) : cfcHom ha = φ := - (cfcHom ha).ext_continuousMap a φ (cfcHom_closedEmbedding ha).continuous hφ₁ <| by + (cfcHom ha).ext_continuousMap a φ (cfcHom_isClosedEmbedding ha).continuous hφ₁ <| by rw [cfcHom_id ha, hφ₂] theorem cfcHom_comp [UniqueContinuousFunctionalCalculus R A] (f : C(spectrum R a, R)) @@ -263,7 +266,7 @@ theorem cfcHom_comp [UniqueContinuousFunctionalCalculus R A] (f : C(spectrum R a (cfcHom ha).comp <| ContinuousMap.compStarAlgHom' R R f' suffices cfcHom (cfcHom_predicate ha f) = φ from DFunLike.congr_fun this.symm g refine cfcHom_eq_of_continuous_of_map_id (cfcHom_predicate ha f) φ ?_ ?_ - · exact (cfcHom_closedEmbedding ha).continuous.comp f'.continuous_comp_left + · exact (cfcHom_isClosedEmbedding ha).continuous.comp f'.continuous_comp_left · simp only [φ, StarAlgHom.comp_apply, ContinuousMap.compStarAlgHom'_apply] congr ext x @@ -279,7 +282,7 @@ noncomputable def cfcL {a : A} (ha : p a) : C(spectrum R a, R) →L[R] A := { cfcHom ha with toFun := cfcHom ha map_smul' := map_smul _ - cont := (cfcHom_closedEmbedding ha).continuous } + cont := (cfcHom_isClosedEmbedding ha).continuous } end cfcL @@ -398,7 +401,7 @@ lemma eqOn_of_cfc_eq_cfc {f g : R → R} {a : A} (h : cfc f a = cfc g a) (hg : ContinuousOn g (spectrum R a) := by cfc_cont_tac) (ha : p a := by cfc_tac) : (spectrum R a).EqOn f g := by rw [cfc_apply f a, cfc_apply g a] at h - have := (cfcHom_closedEmbedding (show p a from ha) (R := R)).inj h + have := (cfcHom_isClosedEmbedding (show p a from ha) (R := R)).inj h intro x hx congrm($(this) ⟨x, hx⟩) diff --git a/Mathlib/Analysis/CStarAlgebra/GelfandDuality.lean b/Mathlib/Analysis/CStarAlgebra/GelfandDuality.lean index af17ea013e3fc..a701ede3a9749 100644 --- a/Mathlib/Analysis/CStarAlgebra/GelfandDuality.lean +++ b/Mathlib/Analysis/CStarAlgebra/GelfandDuality.lean @@ -170,7 +170,7 @@ theorem gelfandTransform_bijective : Function.Bijective (gelfandTransform ℂ A) points in `C(characterSpace ℂ A, ℂ)` and is closed under `star`. -/ have h : rng.topologicalClosure = rng := le_antisymm (StarSubalgebra.topologicalClosure_minimal le_rfl - (gelfandTransform_isometry A).closedEmbedding.isClosed_range) + (gelfandTransform_isometry A).isClosedEmbedding.isClosed_range) (StarSubalgebra.le_topologicalClosure _) refine h ▸ ContinuousMap.starSubalgebra_topologicalClosure_eq_top_of_separatesPoints _ (fun _ _ => ?_) diff --git a/Mathlib/Analysis/Calculus/LineDeriv/IntegrationByParts.lean b/Mathlib/Analysis/Calculus/LineDeriv/IntegrationByParts.lean index 81863eb235bbc..21a62a2741891 100644 --- a/Mathlib/Analysis/Calculus/LineDeriv/IntegrationByParts.lean +++ b/Mathlib/Analysis/Calculus/LineDeriv/IntegrationByParts.lean @@ -130,7 +130,7 @@ theorem integral_bilinear_hasLineDerivAt_right_eq_neg_left_of_integrable -∫ (x : E' × ℝ), (B (f' (L.symm x))) (g (L.symm x)) ∂ν by have : μ = Measure.map L.symm ν := by simp [Measure.map_map L.symm.continuous.measurable L.continuous.measurable] - have hL : ClosedEmbedding L.symm := L.symm.toHomeomorph.closedEmbedding + have hL : IsClosedEmbedding L.symm := L.symm.toHomeomorph.isClosedEmbedding simpa [this, hL.integral_map] using H have L_emb : MeasurableEmbedding L := L.toHomeomorph.measurableEmbedding apply integral_bilinear_hasLineDerivAt_right_eq_neg_left_of_integrable_aux2 diff --git a/Mathlib/Analysis/Convex/Intrinsic.lean b/Mathlib/Analysis/Convex/Intrinsic.lean index 82cc1b1b0da2d..1706893f1b382 100644 --- a/Mathlib/Analysis/Convex/Intrinsic.lean +++ b/Mathlib/Analysis/Convex/Intrinsic.lean @@ -176,11 +176,11 @@ theorem intrinsicFrontier_union_intrinsicInterior (s : Set P) : theorem isClosed_intrinsicClosure (hs : IsClosed (affineSpan 𝕜 s : Set P)) : IsClosed (intrinsicClosure 𝕜 s) := - (closedEmbedding_subtype_val hs).isClosedMap _ isClosed_closure + hs.isClosedEmbedding_subtypeVal.isClosedMap _ isClosed_closure theorem isClosed_intrinsicFrontier (hs : IsClosed (affineSpan 𝕜 s : Set P)) : IsClosed (intrinsicFrontier 𝕜 s) := - (closedEmbedding_subtype_val hs).isClosedMap _ isClosed_frontier + hs.isClosedEmbedding_subtypeVal.isClosedMap _ isClosed_frontier @[simp] theorem affineSpan_intrinsicClosure (s : Set P) : diff --git a/Mathlib/Analysis/Convolution.lean b/Mathlib/Analysis/Convolution.lean index 55806cb2f24fa..00f566fb58ae2 100644 --- a/Mathlib/Analysis/Convolution.lean +++ b/Mathlib/Analysis/Convolution.lean @@ -1239,7 +1239,7 @@ theorem contDiffOn_convolution_right_with_param {f : G → E} {n : ℕ∞} (L : ((ContinuousLinearEquiv.arrowCongr isoE' isoF).symm : (E' →L[𝕜] F) →L[𝕜] eE' →L[𝕜] eF) L let R := fun q : eP × eG => (ef ⋆[eL, eμ] eg q.1) q.2 have R_contdiff : ContDiffOn 𝕜 n R ((isoP ⁻¹' s) ×ˢ univ) := by - have hek : IsCompact (isoG ⁻¹' k) := isoG.toHomeomorph.closedEmbedding.isCompact_preimage hk + have hek : IsCompact (isoG ⁻¹' k) := isoG.toHomeomorph.isClosedEmbedding.isCompact_preimage hk have hes : IsOpen (isoP ⁻¹' s) := isoP.continuous.isOpen_preimage _ hs refine contDiffOn_convolution_right_with_param_aux eL hes hek ?_ ?_ ?_ · intro p x hp hx @@ -1264,9 +1264,9 @@ theorem contDiffOn_convolution_right_with_param {f : G → E} {n : ℕ∞} (L : simp only [LinearIsometryEquiv.coe_coe, (· ∘ ·), ContinuousLinearEquiv.prod_symm, ContinuousLinearEquiv.prod_apply] simp only [R, convolution, coe_comp', ContinuousLinearEquiv.coe_coe, (· ∘ ·)] - rw [ClosedEmbedding.integral_map, ← isoF.integral_comp_comm] + rw [IsClosedEmbedding.integral_map, ← isoF.integral_comp_comm] · rfl - · exact isoG.symm.toHomeomorph.closedEmbedding + · exact isoG.symm.toHomeomorph.isClosedEmbedding simp_rw [this] at A exact A diff --git a/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean b/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean index 31f90aefe8f74..39402193d24a1 100644 --- a/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean +++ b/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean @@ -334,7 +334,7 @@ theorem lintegral_pow_le_pow_lintegral_fderiv_aux [Fintype ι] _ ≤ ∫⁻ xᵢ in Iic (x i), ‖deriv (u ∘ update x i) xᵢ‖₊ := by apply le_trans (by simp) (HasCompactSupport.ennnorm_le_lintegral_Ici_deriv _ _ _) · exact hu.comp (by convert contDiff_update 1 x i) - · exact h2u.comp_closedEmbedding (closedEmbedding_update x i) + · exact h2u.comp_isClosedEmbedding (isClosedEmbedding_update x i) _ ≤ ∫⁻ xᵢ, (‖fderiv ℝ u (update x i xᵢ)‖₊ : ℝ≥0∞) := ?_ gcongr with y; swap · exact Measure.restrict_le_self diff --git a/Mathlib/Analysis/Normed/Algebra/Spectrum.lean b/Mathlib/Analysis/Normed/Algebra/Spectrum.lean index 04202dd932598..942cd2eeae3a2 100644 --- a/Mathlib/Analysis/Normed/Algebra/Spectrum.lean +++ b/Mathlib/Analysis/Normed/Algebra/Spectrum.lean @@ -133,7 +133,7 @@ instance instCompactSpaceNNReal {A : Type*} [NormedRing A] [NormedAlgebra ℝ A] (a : A) [CompactSpace (spectrum ℝ a)] : CompactSpace (spectrum ℝ≥0 a) := by rw [← isCompact_iff_compactSpace] at * rw [← preimage_algebraMap ℝ] - exact closedEmbedding_subtype_val isClosed_nonneg |>.isCompact_preimage <| by assumption + exact isClosed_nonneg.isClosedEmbedding_subtypeVal.isCompact_preimage <| by assumption section QuasispectrumCompact @@ -154,7 +154,7 @@ instance _root_.quasispectrum.instCompactSpaceNNReal [NormedSpace ℝ B] [IsScal CompactSpace (quasispectrum ℝ≥0 a) := by rw [← isCompact_iff_compactSpace] at * rw [← quasispectrum.preimage_algebraMap ℝ] - exact closedEmbedding_subtype_val isClosed_nonneg |>.isCompact_preimage <| by assumption + exact isClosed_nonneg.isClosedEmbedding_subtypeVal.isCompact_preimage <| by assumption end QuasispectrumCompact diff --git a/Mathlib/Analysis/Normed/Module/Basic.lean b/Mathlib/Analysis/Normed/Module/Basic.lean index d1bb26f18316d..09973f5989c6d 100644 --- a/Mathlib/Analysis/Normed/Module/Basic.lean +++ b/Mathlib/Analysis/Normed/Module/Basic.lean @@ -228,8 +228,8 @@ protected theorem NormedSpace.noncompactSpace : NoncompactSpace E := by exact ⟨fun h ↦ NormedSpace.unbounded_univ 𝕜 E h.isBounded⟩ · push_neg at H rcases exists_ne (0 : E) with ⟨x, hx⟩ - suffices ClosedEmbedding (Infinite.natEmbedding 𝕜 · • x) from this.noncompactSpace - refine closedEmbedding_of_pairwise_le_dist (norm_pos_iff.2 hx) fun k n hne ↦ ?_ + suffices IsClosedEmbedding (Infinite.natEmbedding 𝕜 · • x) from this.noncompactSpace + refine isClosedEmbedding_of_pairwise_le_dist (norm_pos_iff.2 hx) fun k n hne ↦ ?_ simp only [dist_eq_norm, ← sub_smul, norm_smul] rw [H, one_mul] rwa [sub_ne_zero, (Embedding.injective _).ne_iff] diff --git a/Mathlib/Analysis/Normed/Module/FiniteDimension.lean b/Mathlib/Analysis/Normed/Module/FiniteDimension.lean index 4735888f210dd..34b802ea77856 100644 --- a/Mathlib/Analysis/Normed/Module/FiniteDimension.lean +++ b/Mathlib/Analysis/Normed/Module/FiniteDimension.lean @@ -519,8 +519,8 @@ lemma ProperSpace.of_locallyCompact_module [Nontrivial E] [LocallyCompactSpace E have : LocallyCompactSpace 𝕜 := by obtain ⟨v, hv⟩ : ∃ v : E, v ≠ 0 := exists_ne 0 let L : 𝕜 → E := fun t ↦ t • v - have : ClosedEmbedding L := closedEmbedding_smul_left hv - apply ClosedEmbedding.locallyCompactSpace this + have : IsClosedEmbedding L := isClosedEmbedding_smul_left hv + apply IsClosedEmbedding.locallyCompactSpace this .of_locallyCompactSpace 𝕜 @[deprecated (since := "2024-01-31")] diff --git a/Mathlib/Analysis/Normed/Operator/Compact.lean b/Mathlib/Analysis/Normed/Operator/Compact.lean index 633508c11661e..da69e15c03410 100644 --- a/Mathlib/Analysis/Normed/Operator/Compact.lean +++ b/Mathlib/Analysis/Normed/Operator/Compact.lean @@ -255,7 +255,7 @@ theorem IsCompactOperator.codRestrict {f : M₁ → M₂} (hf : IsCompactOperato (hV : ∀ x, f x ∈ V) (h_closed : IsClosed (V : Set M₂)) : IsCompactOperator (Set.codRestrict f V hV) := let ⟨_, hK, hKf⟩ := hf - ⟨_, (closedEmbedding_subtype_val h_closed).isCompact_preimage hK, hKf⟩ + ⟨_, h_closed.isClosedEmbedding_subtypeVal.isCompact_preimage hK, hKf⟩ end CodRestrict diff --git a/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/ExpLog.lean b/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/ExpLog.lean index 8c7bd3f605070..65c353f847076 100644 --- a/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/ExpLog.lean +++ b/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/ExpLog.lean @@ -64,7 +64,7 @@ variable {𝕜 : Type*} {A : Type*} [RCLike 𝕜] {p : A → Prop} [NormedRing A lemma exp_eq_normedSpace_exp {a : A} (ha : p a := by cfc_tac) : cfc (exp 𝕜 : 𝕜 → 𝕜) a = exp 𝕜 a := by conv_rhs => rw [← cfc_id 𝕜 a ha, cfc_apply id a ha] - have h := (cfcHom_closedEmbedding (R := 𝕜) (show p a from ha)).continuous + have h := (cfcHom_isClosedEmbedding (R := 𝕜) (show p a from ha)).continuous have _ : ContinuousOn (exp 𝕜) (spectrum 𝕜 a) := exp_continuous.continuousOn simp_rw [← map_exp 𝕜 _ h, cfc_apply (exp 𝕜) a ha] congr 1 diff --git a/Mathlib/Analysis/SpecialFunctions/Log/ENNRealLogExp.lean b/Mathlib/Analysis/SpecialFunctions/Log/ENNRealLogExp.lean index 27f458ecd81d8..e257ffe48d381 100644 --- a/Mathlib/Analysis/SpecialFunctions/Log/ENNRealLogExp.lean +++ b/Mathlib/Analysis/SpecialFunctions/Log/ENNRealLogExp.lean @@ -135,5 +135,5 @@ end Measurability end ENNReal instance : PolishSpace EReal := - ClosedEmbedding.polishSpace ⟨ENNReal.logOrderIso.symm.toHomeomorph.embedding, + IsClosedEmbedding.polishSpace ⟨ENNReal.logOrderIso.symm.toHomeomorph.embedding, ENNReal.logOrderIso.symm.toHomeomorph.range_coe ▸ isClosed_univ⟩ diff --git a/Mathlib/CategoryTheory/Galois/Topology.lean b/Mathlib/CategoryTheory/Galois/Topology.lean index 06ad6569ca157..f0bb5bf4ec1bb 100644 --- a/Mathlib/CategoryTheory/Galois/Topology.lean +++ b/Mathlib/CategoryTheory/Galois/Topology.lean @@ -79,26 +79,29 @@ lemma autEmbedding_range_isClosed : IsClosed (Set.range (autEmbedding F)) := by · fun_prop · fun_prop -lemma autEmbedding_closedEmbedding : ClosedEmbedding (autEmbedding F) where +lemma autEmbedding_isClosedEmbedding : IsClosedEmbedding (autEmbedding F) where induced := rfl inj := autEmbedding_injective F isClosed_range := autEmbedding_range_isClosed F -instance : CompactSpace (Aut F) := ClosedEmbedding.compactSpace (autEmbedding_closedEmbedding F) +@[deprecated (since := "2024-10-20")] +alias autEmbedding_closedEmbedding := autEmbedding_isClosedEmbedding + +instance : CompactSpace (Aut F) := IsClosedEmbedding.compactSpace (autEmbedding_isClosedEmbedding F) instance : T2Space (Aut F) := T2Space.of_injective_continuous (autEmbedding_injective F) continuous_induced_dom instance : TotallyDisconnectedSpace (Aut F) := - (Embedding.isTotallyDisconnected_range (autEmbedding_closedEmbedding F).embedding).mp + (Embedding.isTotallyDisconnected_range (autEmbedding_isClosedEmbedding F).embedding).mp (isTotallyDisconnected_of_totallyDisconnectedSpace _) instance : ContinuousMul (Aut F) := Inducing.continuousMul (autEmbedding F) - (autEmbedding_closedEmbedding F).toInducing + (autEmbedding_isClosedEmbedding F).toInducing instance : ContinuousInv (Aut F) := - Inducing.continuousInv (autEmbedding_closedEmbedding F).toInducing (fun _ ↦ rfl) + Inducing.continuousInv (autEmbedding_isClosedEmbedding F).toInducing (fun _ ↦ rfl) instance : TopologicalGroup (Aut F) := ⟨⟩ diff --git a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean index ac7279ed0be9c..85e1eada2814c 100644 --- a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean +++ b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean @@ -280,19 +280,22 @@ protected theorem image_eq (s : Set H) : I '' s = I.symm ⁻¹' s ∩ range I := · rw [I.source_eq]; exact subset_univ _ · rw [inter_comm, I.target_eq, I.toPartialEquiv_coe_symm] -protected theorem closedEmbedding : ClosedEmbedding I := - I.leftInverse.closedEmbedding I.continuous_symm I.continuous +theorem isClosedEmbedding : IsClosedEmbedding I := + I.leftInverse.isClosedEmbedding I.continuous_symm I.continuous + +@[deprecated (since := "2024-10-20")] +alias closedEmbedding := isClosedEmbedding theorem isClosed_range : IsClosed (range I) := - I.closedEmbedding.isClosed_range + I.isClosedEmbedding.isClosed_range @[deprecated (since := "2024-03-17")] alias closed_range := isClosed_range theorem map_nhds_eq (x : H) : map I (𝓝 x) = 𝓝[range I] I x := - I.closedEmbedding.toEmbedding.map_nhds_eq x + I.isClosedEmbedding.toEmbedding.map_nhds_eq x theorem map_nhdsWithin_eq (s : Set H) (x : H) : map I (𝓝[s] x) = 𝓝[I '' s] I x := - I.closedEmbedding.toEmbedding.map_nhdsWithin_eq s x + I.isClosedEmbedding.toEmbedding.map_nhdsWithin_eq s x theorem image_mem_nhdsWithin {x : H} {s : Set H} (hs : s ∈ 𝓝 x) : I '' s ∈ 𝓝[range I] I x := I.map_nhds_eq x ▸ image_mem_map hs @@ -348,7 +351,7 @@ open TopologicalSpace protected theorem secondCountableTopology [SecondCountableTopology E] (I : ModelWithCorners 𝕜 E H) : SecondCountableTopology H := - I.closedEmbedding.toEmbedding.secondCountableTopology + I.isClosedEmbedding.toEmbedding.secondCountableTopology end ModelWithCorners diff --git a/Mathlib/Geometry/Manifold/WhitneyEmbedding.lean b/Mathlib/Geometry/Manifold/WhitneyEmbedding.lean index 3e3f2e4bb92fd..4e56b5f717089 100644 --- a/Mathlib/Geometry/Manifold/WhitneyEmbedding.lean +++ b/Mathlib/Geometry/Manifold/WhitneyEmbedding.lean @@ -128,9 +128,9 @@ supports of bump functions, then for some `n` it can be embedded into the `n`-di Euclidean space. -/ theorem exists_embedding_euclidean_of_compact [T2Space M] [CompactSpace M] : ∃ (n : ℕ) (e : M → EuclideanSpace ℝ (Fin n)), - Smooth I (𝓡 n) e ∧ ClosedEmbedding e ∧ ∀ x : M, Injective (mfderiv I (𝓡 n) e x) := by + Smooth I (𝓡 n) e ∧ IsClosedEmbedding e ∧ ∀ x : M, Injective (mfderiv I (𝓡 n) e x) := by rcases SmoothBumpCovering.exists_isSubordinate I isClosed_univ fun (x : M) _ => univ_mem with ⟨ι, f, -⟩ haveI := f.fintype rcases f.exists_immersion_euclidean with ⟨n, e, hsmooth, hinj, hinj_mfderiv⟩ - exact ⟨n, e, hsmooth, hsmooth.continuous.closedEmbedding hinj, hinj_mfderiv⟩ + exact ⟨n, e, hsmooth, hsmooth.continuous.isClosedEmbedding hinj, hinj_mfderiv⟩ diff --git a/Mathlib/LinearAlgebra/Matrix/HermitianFunctionalCalculus.lean b/Mathlib/LinearAlgebra/Matrix/HermitianFunctionalCalculus.lean index 5c9c8e5757541..a1b13157f47c4 100644 --- a/Mathlib/LinearAlgebra/Matrix/HermitianFunctionalCalculus.lean +++ b/Mathlib/LinearAlgebra/Matrix/HermitianFunctionalCalculus.lean @@ -87,10 +87,10 @@ noncomputable def cfcAux : C(spectrum ℝ A, ℝ) →⋆ₐ[ℝ] (Matrix n n ext simp -lemma closedEmbedding_cfcAux : ClosedEmbedding hA.cfcAux := by +lemma isClosedEmbedding_cfcAux : IsClosedEmbedding hA.cfcAux := by have h0 : FiniteDimensional ℝ C(spectrum ℝ A, ℝ) := FiniteDimensional.of_injective (ContinuousMap.coeFnLinearMap ℝ (M := ℝ)) DFunLike.coe_injective - refine LinearMap.closedEmbedding_of_injective (𝕜 := ℝ) (E := C(spectrum ℝ A, ℝ)) + refine LinearMap.isClosedEmbedding_of_injective (𝕜 := ℝ) (E := C(spectrum ℝ A, ℝ)) (F := Matrix n n 𝕜) (f := hA.cfcAux) <| LinearMap.ker_eq_bot'.mpr fun f hf ↦ ?_ have h2 : diagonal (RCLike.ofReal ∘ f ∘ fun i ↦ ⟨hA.eigenvalues i, hA.eigenvalues_mem_spectrum_real i⟩) @@ -107,6 +107,9 @@ lemma closedEmbedding_cfcAux : ClosedEmbedding hA.cfcAux := by have := (diagonal_eq_diagonal_iff).mp h2 refine RCLike.ofReal_eq_zero.mp (this i) +@[deprecated (since := "2024-10-20")] +alias closedEmbedding_cfcAux := isClosedEmbedding_cfcAux + lemma cfcAux_id : hA.cfcAux (.restrict (spectrum ℝ A) (.id ℝ)) = A := by conv_rhs => rw [hA.spectral_theorem] congr! @@ -117,7 +120,7 @@ instance instContinuousFunctionalCalculus : ContinuousFunctionalCalculus ℝ (IsSelfAdjoint : Matrix n n 𝕜 → Prop) where exists_cfc_of_predicate a ha := by replace ha : IsHermitian a := ha - refine ⟨ha.cfcAux, ha.closedEmbedding_cfcAux, ha.cfcAux_id, fun f ↦ ?map_spec, + refine ⟨ha.cfcAux, ha.isClosedEmbedding_cfcAux, ha.cfcAux_id, fun f ↦ ?map_spec, fun f ↦ ?hermitian⟩ case map_spec => apply Set.eq_of_subset_of_subset @@ -159,7 +162,7 @@ protected noncomputable def cfc (f : ℝ → ℝ) : Matrix n n 𝕜 := lemma cfc_eq (f : ℝ → ℝ) : cfc f A = hA.cfc f := by have hA' : IsSelfAdjoint A := hA - have := cfcHom_eq_of_continuous_of_map_id hA' hA.cfcAux hA.closedEmbedding_cfcAux.continuous + have := cfcHom_eq_of_continuous_of_map_id hA' hA.cfcAux hA.isClosedEmbedding_cfcAux.continuous hA.cfcAux_id rw [cfc_apply f A hA' (by rw [continuousOn_iff_continuous_restrict]; fun_prop), this] simp only [cfcAux_apply, ContinuousMap.coe_mk, Function.comp_def, Set.restrict_apply, diff --git a/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean b/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean index 601ec698d26dd..c5518f43dd87a 100644 --- a/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean +++ b/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean @@ -472,9 +472,12 @@ is ae-measurable. -/ theorem Continuous.aemeasurable {f : α → γ} (h : Continuous f) {μ : Measure α} : AEMeasurable f μ := h.measurable.aemeasurable -theorem ClosedEmbedding.measurable {f : α → γ} (hf : ClosedEmbedding f) : Measurable f := +theorem IsClosedEmbedding.measurable {f : α → γ} (hf : IsClosedEmbedding f) : Measurable f := hf.continuous.measurable +@[deprecated (since := "2024-10-20")] +alias ClosedEmbedding.measurable := IsClosedEmbedding.measurable + /-- If a function is defined piecewise in terms of functions which are continuous on their respective pieces, then it is measurable. -/ theorem ContinuousOn.measurable_piecewise {f g : α → γ} {s : Set α} [∀ j : α, Decidable (j ∈ s)] @@ -632,10 +635,13 @@ protected theorem Embedding.measurableEmbedding {f : α → β} (h₁ : Embeddin (((↑) : range f → β) ∘ (Homeomorph.ofEmbedding f h₁).toMeasurableEquiv) from (MeasurableEmbedding.subtype_coe h₂).comp (MeasurableEquiv.measurableEmbedding _) -protected theorem ClosedEmbedding.measurableEmbedding {f : α → β} (h : ClosedEmbedding f) : +protected theorem IsClosedEmbedding.measurableEmbedding {f : α → β} (h : IsClosedEmbedding f) : MeasurableEmbedding f := h.toEmbedding.measurableEmbedding h.isClosed_range.measurableSet +@[deprecated (since := "2024-10-20")] +alias ClosedEmbedding.measurableEmbedding := IsClosedEmbedding.measurableEmbedding + protected theorem IsOpenEmbedding.measurableEmbedding {f : α → β} (h : IsOpenEmbedding f) : MeasurableEmbedding f := h.toEmbedding.measurableEmbedding h.isOpen_range.measurableSet diff --git a/Mathlib/MeasureTheory/Constructions/BorelSpace/ContinuousLinearMap.lean b/Mathlib/MeasureTheory/Constructions/BorelSpace/ContinuousLinearMap.lean index 32e8e945358ec..c32b095c05312 100644 --- a/Mathlib/MeasureTheory/Constructions/BorelSpace/ContinuousLinearMap.lean +++ b/Mathlib/MeasureTheory/Constructions/BorelSpace/ContinuousLinearMap.lean @@ -90,10 +90,10 @@ variable [BorelSpace 𝕜] {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 theorem measurable_smul_const {f : α → 𝕜} {c : E} (hc : c ≠ 0) : (Measurable fun x => f x • c) ↔ Measurable f := - (closedEmbedding_smul_left hc).measurableEmbedding.measurable_comp_iff + (isClosedEmbedding_smul_left hc).measurableEmbedding.measurable_comp_iff theorem aemeasurable_smul_const {f : α → 𝕜} {μ : Measure α} {c : E} (hc : c ≠ 0) : AEMeasurable (fun x => f x • c) μ ↔ AEMeasurable f μ := - (closedEmbedding_smul_left hc).measurableEmbedding.aemeasurable_comp_iff + (isClosedEmbedding_smul_left hc).measurableEmbedding.aemeasurable_comp_iff end NormedSpace diff --git a/Mathlib/MeasureTheory/Constructions/Polish/EmbeddingReal.lean b/Mathlib/MeasureTheory/Constructions/Polish/EmbeddingReal.lean index 67a330bbb5bb7..49a72519b60d3 100644 --- a/Mathlib/MeasureTheory/Constructions/Polish/EmbeddingReal.lean +++ b/Mathlib/MeasureTheory/Constructions/Polish/EmbeddingReal.lean @@ -24,7 +24,7 @@ theorem exists_nat_measurableEquiv_range_coe_fin_of_finite [Finite α] : theorem measurableEquiv_range_coe_nat_of_infinite_of_countable [Infinite α] [Countable α] : Nonempty (α ≃ᵐ range ((↑) : ℕ → ℝ)) := by have : PolishSpace (range ((↑) : ℕ → ℝ)) := - Nat.closedEmbedding_coe_real.isClosedMap.isClosed_range.polishSpace + Nat.isClosedEmbedding_coe_real.isClosedMap.isClosed_range.polishSpace refine ⟨PolishSpace.Equiv.measurableEquiv ?_⟩ refine (nonempty_equiv_of_countable.some : α ≃ ℕ).trans ?_ exact Equiv.ofInjective ((↑) : ℕ → ℝ) Nat.cast_injective diff --git a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean index 7d3943e021a7c..012d894c0ae6f 100644 --- a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean +++ b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean @@ -691,7 +691,7 @@ theorem _root_.Embedding.comp_stronglyMeasurable_iff {m : MeasurableSpace α} [T ⟨fun H => stronglyMeasurable_iff_measurable_separable.2 ⟨?_, ?_⟩, fun H => hg.continuous.comp_stronglyMeasurable H⟩ · let G : β → range g := rangeFactorization g - have hG : ClosedEmbedding G := + have hG : IsClosedEmbedding G := { hg.codRestrict _ _ with isClosed_range := by rw [surjective_onto_range.range_eq] @@ -1513,7 +1513,7 @@ theorem _root_.Embedding.aestronglyMeasurable_comp_iff [PseudoMetrizableSpace β ⟨fun H => aestronglyMeasurable_iff_aemeasurable_separable.2 ⟨?_, ?_⟩, fun H => hg.continuous.comp_aestronglyMeasurable H⟩ · let G : β → range g := rangeFactorization g - have hG : ClosedEmbedding G := + have hG : IsClosedEmbedding G := { hg.codRestrict _ _ with isClosed_range := by rw [surjective_onto_range.range_eq]; exact isClosed_univ } have : AEMeasurable (G ∘ f) μ := AEMeasurable.subtype_mk H.aemeasurable diff --git a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Lemmas.lean b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Lemmas.lean index 4251891161db2..95fbcc33ea7cf 100644 --- a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Lemmas.lean +++ b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Lemmas.lean @@ -46,7 +46,7 @@ variable {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] theorem aestronglyMeasurable_smul_const_iff {f : α → 𝕜} {c : E} (hc : c ≠ 0) : AEStronglyMeasurable (fun x => f x • c) μ ↔ AEStronglyMeasurable f μ := - (closedEmbedding_smul_left hc).toEmbedding.aestronglyMeasurable_comp_iff + (isClosedEmbedding_smul_left hc).toEmbedding.aestronglyMeasurable_comp_iff end NormedSpace diff --git a/Mathlib/MeasureTheory/Group/Measure.lean b/Mathlib/MeasureTheory/Group/Measure.lean index e058bb5e8e9af..49764e92934c6 100644 --- a/Mathlib/MeasureTheory/Group/Measure.lean +++ b/Mathlib/MeasureTheory/Group/Measure.lean @@ -775,7 +775,7 @@ nonrec theorem _root_.MulEquiv.isHaarMeasure_map [BorelSpace G] [TopologicalGrou [TopologicalGroup H] (e : G ≃* H) (he : Continuous e) (hesymm : Continuous e.symm) : IsHaarMeasure (Measure.map e μ) := let f : G ≃ₜ H := .mk e - isHaarMeasure_map μ (e : G →* H) he e.surjective f.closedEmbedding.tendsto_cocompact + isHaarMeasure_map μ (e : G →* H) he e.surjective f.isClosedEmbedding.tendsto_cocompact /-- A convenience wrapper for MeasureTheory.Measure.isAddHaarMeasure_map`. -/ instance _root_.ContinuousLinearEquiv.isAddHaarMeasure_map diff --git a/Mathlib/MeasureTheory/Integral/Bochner.lean b/Mathlib/MeasureTheory/Integral/Bochner.lean index b497a13fc272b..3c5f56fce5d8b 100644 --- a/Mathlib/MeasureTheory/Integral/Bochner.lean +++ b/Mathlib/MeasureTheory/Integral/Bochner.lean @@ -1574,11 +1574,14 @@ theorem _root_.MeasurableEmbedding.integral_map {β} {_ : MeasurableSpace β} {f · rw [integral_non_aestronglyMeasurable hgm, integral_non_aestronglyMeasurable] exact fun hgf => hgm (hf.aestronglyMeasurable_map_iff.2 hgf) -theorem _root_.ClosedEmbedding.integral_map {β} [TopologicalSpace α] [BorelSpace α] - [TopologicalSpace β] [MeasurableSpace β] [BorelSpace β] {φ : α → β} (hφ : ClosedEmbedding φ) +theorem _root_.IsClosedEmbedding.integral_map {β} [TopologicalSpace α] [BorelSpace α] + [TopologicalSpace β] [MeasurableSpace β] [BorelSpace β] {φ : α → β} (hφ : IsClosedEmbedding φ) (f : β → G) : ∫ y, f y ∂Measure.map φ μ = ∫ x, f (φ x) ∂μ := hφ.measurableEmbedding.integral_map _ +@[deprecated (since := "2024-10-20")] +alias _root_.ClosedEmbedding.integral_map := _root_.IsClosedEmbedding.integral_map + theorem integral_map_equiv {β} [MeasurableSpace β] (e : α ≃ᵐ β) (f : β → G) : ∫ y, f y ∂Measure.map e μ = ∫ x, f (e x) ∂μ := e.measurableEmbedding.integral_map f diff --git a/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean b/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean index 0c443e0792ad6..6b904ee39dabd 100644 --- a/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean +++ b/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean @@ -271,7 +271,7 @@ theorem comp_mul_left (hf : IntervalIntegrable f volume a b) (c : ℝ) : rcases eq_or_ne c 0 with (hc | hc); · rw [hc]; simp rw [intervalIntegrable_iff'] at hf ⊢ have A : MeasurableEmbedding fun x => x * c⁻¹ := - (Homeomorph.mulRight₀ _ (inv_ne_zero hc)).closedEmbedding.measurableEmbedding + (Homeomorph.mulRight₀ _ (inv_ne_zero hc)).isClosedEmbedding.measurableEmbedding rw [← Real.smul_map_volume_mul_right (inv_ne_zero hc), IntegrableOn, Measure.restrict_smul, integrable_smul_measure (by simpa : ENNReal.ofReal |c⁻¹| ≠ 0) ENNReal.ofReal_ne_top, ← IntegrableOn, MeasurableEmbedding.integrableOn_map_iff A] @@ -294,7 +294,7 @@ theorem comp_add_right (hf : IntervalIntegrable f volume a b) (c : ℝ) : · exact IntervalIntegrable.symm (this hf.symm (le_of_not_le h)) rw [intervalIntegrable_iff'] at hf ⊢ have A : MeasurableEmbedding fun x => x + c := - (Homeomorph.addRight c).closedEmbedding.measurableEmbedding + (Homeomorph.addRight c).isClosedEmbedding.measurableEmbedding rw [← map_add_right_eq_self volume c] at hf convert (MeasurableEmbedding.integrableOn_map_iff A).mp hf using 1 rw [preimage_add_const_uIcc] @@ -624,7 +624,7 @@ variable {a b c d : ℝ} (f : ℝ → E) theorem integral_comp_mul_right (hc : c ≠ 0) : (∫ x in a..b, f (x * c)) = c⁻¹ • ∫ x in a * c..b * c, f x := by have A : MeasurableEmbedding fun x => x * c := - (Homeomorph.mulRight₀ c hc).closedEmbedding.measurableEmbedding + (Homeomorph.mulRight₀ c hc).isClosedEmbedding.measurableEmbedding conv_rhs => rw [← Real.smul_map_volume_mul_right hc] simp_rw [integral_smul_measure, intervalIntegral, A.setIntegral_map, ENNReal.toReal_ofReal (abs_nonneg c)] @@ -661,7 +661,7 @@ theorem inv_smul_integral_comp_div (c) : @[simp] theorem integral_comp_add_right (d) : (∫ x in a..b, f (x + d)) = ∫ x in a + d..b + d, f x := have A : MeasurableEmbedding fun x => x + d := - (Homeomorph.addRight d).closedEmbedding.measurableEmbedding + (Homeomorph.addRight d).isClosedEmbedding.measurableEmbedding calc (∫ x in a..b, f (x + d)) = ∫ x in a + d..b + d, f x ∂Measure.map (fun x => x + d) volume := by simp [intervalIntegral, A.setIntegral_map] diff --git a/Mathlib/MeasureTheory/Integral/SetIntegral.lean b/Mathlib/MeasureTheory/Integral/SetIntegral.lean index 76b9b535d4c5c..6376b28c74713 100644 --- a/Mathlib/MeasureTheory/Integral/SetIntegral.lean +++ b/Mathlib/MeasureTheory/Integral/SetIntegral.lean @@ -558,13 +558,19 @@ theorem _root_.MeasurableEmbedding.setIntegral_map {Y} {_ : MeasurableSpace Y} { @[deprecated (since := "2024-04-17")] alias _root_.MeasurableEmbedding.set_integral_map := _root_.MeasurableEmbedding.setIntegral_map -theorem _root_.ClosedEmbedding.setIntegral_map [TopologicalSpace X] [BorelSpace X] {Y} +theorem _root_.IsClosedEmbedding.setIntegral_map [TopologicalSpace X] [BorelSpace X] {Y} [MeasurableSpace Y] [TopologicalSpace Y] [BorelSpace Y] {g : X → Y} {f : Y → E} (s : Set Y) - (hg : ClosedEmbedding g) : ∫ y in s, f y ∂Measure.map g μ = ∫ x in g ⁻¹' s, f (g x) ∂μ := + (hg : IsClosedEmbedding g) : ∫ y in s, f y ∂Measure.map g μ = ∫ x in g ⁻¹' s, f (g x) ∂μ := hg.measurableEmbedding.setIntegral_map _ _ +@[deprecated (since := "2024-10-20")] +alias _root_.ClosedEmbedding.setIntegral_map := IsClosedEmbedding.setIntegral_map + @[deprecated (since := "2024-04-17")] -alias _root_.ClosedEmbedding.set_integral_map := _root_.ClosedEmbedding.setIntegral_map +alias _root_.IsClosedEmbedding.set_integral_map := _root_.IsClosedEmbedding.setIntegral_map + +@[deprecated (since := "2024-10-20")] +alias _root_.ClosedEmbedding.set_integral_map := IsClosedEmbedding.set_integral_map theorem MeasurePreserving.setIntegral_preimage_emb {Y} {_ : MeasurableSpace Y} {f : X → Y} {ν} (h₁ : MeasurePreserving f μ ν) (h₂ : MeasurableEmbedding f) (g : Y → E) (s : Set Y) : diff --git a/Mathlib/MeasureTheory/Measure/Lebesgue/Integral.lean b/Mathlib/MeasureTheory/Measure/Lebesgue/Integral.lean index 1b31285651c1c..e4968335130b1 100644 --- a/Mathlib/MeasureTheory/Measure/Lebesgue/Integral.lean +++ b/Mathlib/MeasureTheory/Measure/Lebesgue/Integral.lean @@ -80,7 +80,7 @@ itself, it does not apply when `f` is more complicated -/ theorem integral_comp_neg_Iic {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] (c : ℝ) (f : ℝ → E) : (∫ x in Iic c, f (-x)) = ∫ x in Ioi (-c), f x := by have A : MeasurableEmbedding fun x : ℝ => -x := - (Homeomorph.neg ℝ).closedEmbedding.measurableEmbedding + (Homeomorph.neg ℝ).isClosedEmbedding.measurableEmbedding have := MeasurableEmbedding.setIntegral_map (μ := volume) A f (Ici (-c)) rw [Measure.map_neg_eq_self (volume : Measure ℝ)] at this simp_rw [← integral_Ici_eq_integral_Ioi, this, neg_preimage, preimage_neg_Ici, neg_neg] diff --git a/Mathlib/NumberTheory/Modular.lean b/Mathlib/NumberTheory/Modular.lean index d27a687df5826..339c0c6fe0f94 100644 --- a/Mathlib/NumberTheory/Modular.lean +++ b/Mathlib/NumberTheory/Modular.lean @@ -113,7 +113,7 @@ attribute [local simp] ContinuousLinearMap.coe_smul theorem tendsto_normSq_coprime_pair : Filter.Tendsto (fun p : Fin 2 → ℤ => normSq ((p 0 : ℂ) * z + p 1)) cofinite atTop := by -- using this instance rather than the automatic `Function.module` makes unification issues in - -- `LinearEquiv.closedEmbedding_of_injective` less bad later in the proof. + -- `LinearEquiv.isClosedEmbedding_of_injective` less bad later in the proof. letI : Module ℝ (Fin 2 → ℝ) := NormedSpace.toModule let π₀ : (Fin 2 → ℝ) →ₗ[ℝ] ℝ := LinearMap.proj 0 let π₁ : (Fin 2 → ℝ) →ₗ[ℝ] ℝ := LinearMap.proj 1 @@ -149,7 +149,7 @@ theorem tendsto_normSq_coprime_pair : rw [f_def, RingHom.map_add, RingHom.map_mul, mul_add, mul_left_comm, mul_conj, conj_ofReal, conj_ofReal, ← ofReal_mul, add_im, ofReal_im, zero_add, inv_mul_eq_iff_eq_mul₀ hz] simp only [ofReal_im, ofReal_re, mul_im, zero_add, mul_zero] - have hf' : ClosedEmbedding f := f.closedEmbedding_of_injective hf + have hf' : IsClosedEmbedding f := f.isClosedEmbedding_of_injective hf have h₂ : Tendsto (fun p : Fin 2 → ℤ => ((↑) : ℤ → ℝ) ∘ p) cofinite (cocompact _) := by convert Tendsto.pi_map_coprodᵢ fun _ => Int.tendsto_coe_cofinite · rw [coprodᵢ_cofinite] @@ -207,8 +207,8 @@ theorem tendsto_lcRow0 {cd : Fin 2 → ℤ} (hcd : IsCoprime (cd 0) (cd 1)) : Tendsto.pi_map_coprodᵢ fun _ : Fin 2 => Int.tendsto_coe_cofinite have hf₁ : Tendsto f₁ cofinite (cocompact _) := cocompact_ℝ_to_cofinite_ℤ_matrix.comp Subtype.coe_injective.tendsto_cofinite - have hf₂ : ClosedEmbedding (lcRow0Extend hcd) := - (lcRow0Extend hcd).toContinuousLinearEquiv.toHomeomorph.closedEmbedding + have hf₂ : IsClosedEmbedding (lcRow0Extend hcd) := + (lcRow0Extend hcd).toContinuousLinearEquiv.toHomeomorph.isClosedEmbedding convert hf₂.tendsto_cocompact.comp (hf₁.comp Subtype.coe_injective.tendsto_cofinite) using 1 ext ⟨g, rfl⟩ i j : 3 fin_cases i <;> [fin_cases j; skip] @@ -258,7 +258,7 @@ theorem tendsto_abs_re_smul {p : Fin 2 → ℤ} (hp : IsCoprime (p 0) (p 1)) : let f := Homeomorph.mulRight₀ _ this let ff := Homeomorph.addRight (((p 1 : ℂ) * z - p 0) / (((p 0 : ℂ) ^ 2 + (p 1 : ℂ) ^ 2) * (p 0 * z + p 1))).re - convert (f.trans ff).closedEmbedding.tendsto_cocompact.comp (tendsto_lcRow0 hp) with _ _ g + convert (f.trans ff).isClosedEmbedding.tendsto_cocompact.comp (tendsto_lcRow0 hp) with _ _ g change ((g : SL(2, ℤ)) • z).re = lcRow0 p ↑(↑g : SL(2, ℝ)) / ((p 0 : ℝ) ^ 2 + (p 1 : ℝ) ^ 2) + diff --git a/Mathlib/Topology/Algebra/Constructions/DomMulAct.lean b/Mathlib/Topology/Algebra/Constructions/DomMulAct.lean index a907bb2550a2a..575ad9565a29c 100644 --- a/Mathlib/Topology/Algebra/Constructions/DomMulAct.lean +++ b/Mathlib/Topology/Algebra/Constructions/DomMulAct.lean @@ -50,9 +50,13 @@ theorem coe_mkHomeomorph_symm : ⇑(mkHomeomorph : M ≃ₜ Mᵈᵐᵃ).symm = m @[to_additive] theorem inducing_mk : Inducing (@mk M) := mkHomeomorph.inducing @[to_additive] theorem embedding_mk : Embedding (@mk M) := mkHomeomorph.embedding @[to_additive] theorem isOpenEmbedding_mk : IsOpenEmbedding (@mk M) := mkHomeomorph.isOpenEmbedding -@[to_additive] theorem closedEmbedding_mk : ClosedEmbedding (@mk M) := mkHomeomorph.closedEmbedding +@[to_additive] theorem isClosedEmbedding_mk : IsClosedEmbedding (@mk M) := + mkHomeomorph.isClosedEmbedding @[to_additive] theorem quotientMap_mk : QuotientMap (@mk M) := mkHomeomorph.quotientMap +@[deprecated (since := "2024-10-20")] +alias closedEmbedding_mk := isClosedEmbedding_mk + @[deprecated (since := "2024-10-18")] alias openEmbedding_mk := isOpenEmbedding_mk @@ -66,7 +70,11 @@ theorem isOpenEmbedding_mk_symm : IsOpenEmbedding (@mk M).symm := mkHomeomorph.s alias openEmbedding_mk_symm := isOpenEmbedding_mk_symm @[to_additive] -theorem closedEmbedding_mk_symm : ClosedEmbedding (@mk M).symm := mkHomeomorph.symm.closedEmbedding +theorem isClosedEmbedding_mk_symm : IsClosedEmbedding (@mk M).symm := + mkHomeomorph.symm.isClosedEmbedding + +@[deprecated (since := "2024-10-20")] +alias closedEmbedding_mk_symm := isClosedEmbedding_mk_symm @[to_additive] theorem quotientMap_mk_symm : QuotientMap (@mk M).symm := mkHomeomorph.symm.quotientMap @@ -76,8 +84,8 @@ theorem quotientMap_mk_symm : QuotientMap (@mk M).symm := mkHomeomorph.symm.quot @[to_additive] instance instT2Space [T2Space M] : T2Space Mᵈᵐᵃ := embedding_mk_symm.t2Space @[to_additive] instance instT25Space [T25Space M] : T25Space Mᵈᵐᵃ := embedding_mk_symm.t25Space @[to_additive] instance instT3Space [T3Space M] : T3Space Mᵈᵐᵃ := embedding_mk_symm.t3Space -@[to_additive] instance instT4Space [T4Space M] : T4Space Mᵈᵐᵃ := closedEmbedding_mk_symm.t4Space -@[to_additive] instance instT5Space [T5Space M] : T5Space Mᵈᵐᵃ := closedEmbedding_mk_symm.t5Space +@[to_additive] instance instT4Space [T4Space M] : T4Space Mᵈᵐᵃ := isClosedEmbedding_mk_symm.t4Space +@[to_additive] instance instT5Space [T5Space M] : T5Space Mᵈᵐᵃ := isClosedEmbedding_mk_symm.t5Space @[to_additive] instance instR0Space [R0Space M] : R0Space Mᵈᵐᵃ := embedding_mk_symm.r0Space @[to_additive] instance instR1Space [R1Space M] : R1Space Mᵈᵐᵃ := embedding_mk_symm.r1Space @@ -86,11 +94,11 @@ theorem quotientMap_mk_symm : QuotientMap (@mk M).symm := mkHomeomorph.symm.quot instance instRegularSpace [RegularSpace M] : RegularSpace Mᵈᵐᵃ := embedding_mk_symm.regularSpace @[to_additive] -instance instNormalSpace [NormalSpace M] : NormalSpace Mᵈᵐᵃ := closedEmbedding_mk_symm.normalSpace +instance instNormalSpace [NormalSpace M] : NormalSpace Mᵈᵐᵃ := isClosedEmbedding_mk_symm.normalSpace @[to_additive] instance instCompletelyNormalSpace [CompletelyNormalSpace M] : CompletelyNormalSpace Mᵈᵐᵃ := - closedEmbedding_mk_symm.completelyNormalSpace + isClosedEmbedding_mk_symm.completelyNormalSpace @[to_additive] instance instDiscreteTopology [DiscreteTopology M] : DiscreteTopology Mᵈᵐᵃ := @@ -119,7 +127,7 @@ instance instLocallyCompactSpace [LocallyCompactSpace M] : LocallyCompactSpace M @[to_additive] instance instWeaklyLocallyCompactSpace [WeaklyLocallyCompactSpace M] : WeaklyLocallyCompactSpace Mᵈᵐᵃ := - closedEmbedding_mk_symm.weaklyLocallyCompactSpace + isClosedEmbedding_mk_symm.weaklyLocallyCompactSpace @[to_additive (attr := simp)] theorem map_mk_nhds (x : M) : map (mk : M → Mᵈᵐᵃ) (𝓝 x) = 𝓝 (mk x) := diff --git a/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean b/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean index a41b781ceaadc..faf9ab016a5e2 100644 --- a/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean +++ b/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean @@ -258,8 +258,8 @@ lemma range_toContinuousMap : exact ⟨{ f with map_one' := h1, map_mul' := hmul }, rfl⟩ @[to_additive] -theorem closedEmbedding_toContinuousMap [ContinuousMul B] [T2Space B] : - ClosedEmbedding (toContinuousMap : ContinuousMonoidHom A B → C(A, B)) where +theorem isClosedEmbedding_toContinuousMap [ContinuousMul B] [T2Space B] : + IsClosedEmbedding (toContinuousMap : ContinuousMonoidHom A B → C(A, B)) where toEmbedding := embedding_toContinuousMap A B isClosed_range := by simp only [range_toContinuousMap, Set.setOf_and, Set.setOf_forall] @@ -268,6 +268,9 @@ theorem closedEmbedding_toContinuousMap [ContinuousMul B] [T2Space B] : exact isClosed_eq (continuous_eval_const (x * y)) <| .mul (continuous_eval_const x) (continuous_eval_const y) +@[deprecated (since := "2024-10-20")] +alias closedEmbedding_toContinuousMap := isClosedEmbedding_toContinuousMap + variable {A B C D E} @[to_additive] diff --git a/Mathlib/Topology/Algebra/Module/Alternating/Topology.lean b/Mathlib/Topology/Algebra/Module/Alternating/Topology.lean index 52bcb32bef151..61776f9e7c9c2 100644 --- a/Mathlib/Topology/Algebra/Module/Alternating/Topology.lean +++ b/Mathlib/Topology/Algebra/Module/Alternating/Topology.lean @@ -158,11 +158,14 @@ theorem hasBasis_nhds_zero : variable [ContinuousSMul 𝕜 E] -lemma closedEmbedding_toContinuousMultilinearMap [T2Space F] : - ClosedEmbedding (toContinuousMultilinearMap : +lemma isClosedEmbedding_toContinuousMultilinearMap [T2Space F] : + IsClosedEmbedding (toContinuousMultilinearMap : (E [⋀^ι]→L[𝕜] F) → ContinuousMultilinearMap 𝕜 (fun _ : ι ↦ E) F) := ⟨embedding_toContinuousMultilinearMap, isClosed_range_toContinuousMultilinearMap⟩ +@[deprecated (since := "2024-10-20")] +alias closedEmbedding_toContinuousMultilinearMap := isClosedEmbedding_toContinuousMultilinearMap + instance instContinuousEvalConst : ContinuousEvalConst (E [⋀^ι]→L[𝕜] F) (ι → E) F := .of_continuous_forget continuous_toContinuousMultilinearMap diff --git a/Mathlib/Topology/Algebra/Module/FiniteDimension.lean b/Mathlib/Topology/Algebra/Module/FiniteDimension.lean index 2cf817161d306..ccbd57b6b210b 100644 --- a/Mathlib/Topology/Algebra/Module/FiniteDimension.lean +++ b/Mathlib/Topology/Algebra/Module/FiniteDimension.lean @@ -518,24 +518,30 @@ theorem Submodule.closed_of_finiteDimensional s.complete_of_finiteDimensional.isClosed /-- An injective linear map with finite-dimensional domain is a closed embedding. -/ -theorem LinearMap.closedEmbedding_of_injective [T2Space E] [FiniteDimensional 𝕜 E] {f : E →ₗ[𝕜] F} - (hf : LinearMap.ker f = ⊥) : ClosedEmbedding f := +theorem LinearMap.isClosedEmbedding_of_injective [T2Space E] [FiniteDimensional 𝕜 E] {f : E →ₗ[𝕜] F} + (hf : LinearMap.ker f = ⊥) : IsClosedEmbedding f := let g := LinearEquiv.ofInjective f (LinearMap.ker_eq_bot.mp hf) { embedding_subtype_val.comp g.toContinuousLinearEquiv.toHomeomorph.embedding with isClosed_range := by haveI := f.finiteDimensional_range simpa [LinearMap.range_coe f] using f.range.closed_of_finiteDimensional } -theorem closedEmbedding_smul_left [T2Space E] {c : E} (hc : c ≠ 0) : - ClosedEmbedding fun x : 𝕜 => x • c := - LinearMap.closedEmbedding_of_injective (LinearMap.ker_toSpanSingleton 𝕜 E hc) +@[deprecated (since := "2024-10-20")] +alias LinearMap.closedEmbedding_of_injective := LinearMap.isClosedEmbedding_of_injective + +theorem isClosedEmbedding_smul_left [T2Space E] {c : E} (hc : c ≠ 0) : + IsClosedEmbedding fun x : 𝕜 => x • c := + LinearMap.isClosedEmbedding_of_injective (LinearMap.ker_toSpanSingleton 𝕜 E hc) + +@[deprecated (since := "2024-10-20")] +alias closedEmbedding_smul_left := isClosedEmbedding_smul_left -- `smul` is a closed map in the first argument. theorem isClosedMap_smul_left [T2Space E] (c : E) : IsClosedMap fun x : 𝕜 => x • c := by by_cases hc : c = 0 · simp_rw [hc, smul_zero] exact isClosedMap_const - · exact (closedEmbedding_smul_left hc).isClosedMap + · exact (isClosedEmbedding_smul_left hc).isClosedMap theorem ContinuousLinearMap.exists_right_inverse_of_surjective [FiniteDimensional 𝕜 F] (f : E →L[𝕜] F) (hf : LinearMap.range f = ⊤) : diff --git a/Mathlib/Topology/Algebra/ProperAction/Basic.lean b/Mathlib/Topology/Algebra/ProperAction/Basic.lean index 92fc69d56d9f4..8b179ad23ab0f 100644 --- a/Mathlib/Topology/Algebra/ProperAction/Basic.lean +++ b/Mathlib/Topology/Algebra/ProperAction/Basic.lean @@ -127,9 +127,7 @@ then this topological space is T2."] theorem t2Space_of_properSMul_of_t2Group [h_proper : ProperSMul G X] [T2Space G] : T2Space X := by let f := fun x : X ↦ ((1 : G), x) have proper_f : IsProperMap f := by - apply isProperMap_of_closedEmbedding - rw [closedEmbedding_iff] - constructor + refine IsClosedEmbedding.isProperMap ⟨?_, ?_⟩ · let g := fun gx : G × X ↦ gx.2 have : Function.LeftInverse g f := fun x ↦ by simp exact this.embedding (by fun_prop) (by fun_prop) @@ -150,18 +148,20 @@ then `H` also acts properly on `X`. -/ @[to_additive "If two groups `H` and `G` act on a topological space `X` such that `G` acts properly and there exists a group homomorphims `H → G` which is a closed embedding compatible with the actions, then `H` also acts properly on `X`."] -theorem properSMul_of_closedEmbedding {H : Type*} [Group H] [MulAction H X] [TopologicalSpace H] - [ProperSMul G X] (f : H →* G) (f_clemb : ClosedEmbedding f) +theorem properSMul_of_isClosedEmbedding {H : Type*} [Group H] [MulAction H X] [TopologicalSpace H] + [ProperSMul G X] (f : H →* G) (f_clemb : IsClosedEmbedding f) (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)) := this.prodMap isProperMap_id + have h : IsProperMap (Prod.map f (fun x : X ↦ x)) := f_clemb.isProperMap.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] exact h.comp <| ProperSMul.isProperMap_smul_pair +@[deprecated (since := "2024-10-20")] +alias properSMul_of_closedEmbedding := properSMul_of_isClosedEmbedding + /-- If `H` is a closed subgroup of `G` and `G` acts properly on X then so does `H`. -/ @[to_additive "If `H` is a closed subgroup of `G` and `G` acts properly on X then so does `H`."] instance {H : Subgroup G} [ProperSMul G X] [H_closed : IsClosed (H : Set G)] : ProperSMul H X := - properSMul_of_closedEmbedding H.subtype H_closed.closedEmbedding_subtype_val fun _ _ ↦ rfl + properSMul_of_isClosedEmbedding H.subtype H_closed.isClosedEmbedding_subtypeVal fun _ _ ↦ rfl diff --git a/Mathlib/Topology/Algebra/StarSubalgebra.lean b/Mathlib/Topology/Algebra/StarSubalgebra.lean index a8dd137d03c23..295e2416e5169 100644 --- a/Mathlib/Topology/Algebra/StarSubalgebra.lean +++ b/Mathlib/Topology/Algebra/StarSubalgebra.lean @@ -37,9 +37,9 @@ theorem embedding_inclusion {S₁ S₂ : StarSubalgebra R A} (h : S₁ ≤ S₂) { induced := Eq.symm induced_compose inj := Subtype.map_injective h Function.injective_id } -/-- The `StarSubalgebra.inclusion` of a closed star subalgebra is a `ClosedEmbedding`. -/ -theorem closedEmbedding_inclusion {S₁ S₂ : StarSubalgebra R A} (h : S₁ ≤ S₂) - (hS₁ : IsClosed (S₁ : Set A)) : ClosedEmbedding (inclusion h) := +/-- The `StarSubalgebra.inclusion` of a closed star subalgebra is a `IsClosedEmbedding`. -/ +theorem isClosedEmbedding_inclusion {S₁ S₂ : StarSubalgebra R A} (h : S₁ ≤ S₂) + (hS₁ : IsClosed (S₁ : Set A)) : IsClosedEmbedding (inclusion h) := { embedding_inclusion h with isClosed_range := isClosed_induced_iff.2 ⟨S₁, hS₁, by @@ -48,6 +48,9 @@ theorem closedEmbedding_inclusion {S₁ S₂ : StarSubalgebra R A} (h : S₁ ≤ · intro _ h' apply h h' ⟩ } +@[deprecated (since := "2024-10-20")] +alias closedEmbedding_inclusion := isClosedEmbedding_inclusion + variable [TopologicalSemiring A] [ContinuousStar A] variable [TopologicalSpace B] [Semiring B] [Algebra R B] [StarRing B] @@ -99,7 +102,7 @@ theorem map_topologicalClosure_le [StarModule R B] [TopologicalSemiring B] [Cont image_closure_subset_closure_image hφ theorem topologicalClosure_map [StarModule R B] [TopologicalSemiring B] [ContinuousStar B] - (s : StarSubalgebra R A) (φ : A →⋆ₐ[R] B) (hφ : ClosedEmbedding φ) : + (s : StarSubalgebra R A) (φ : A →⋆ₐ[R] B) (hφ : IsClosedEmbedding φ) : (map φ s).topologicalClosure = map φ s.topologicalClosure := SetLike.coe_injective <| hφ.closure_image_eq _ @@ -205,8 +208,8 @@ theorem le_of_isClosed_of_mem {S : StarSubalgebra R A} (hS : IsClosed (S : Set A (hx : x ∈ S) : elementalStarAlgebra R x ≤ S := topologicalClosure_minimal (adjoin_le <| Set.singleton_subset_iff.2 hx) hS -/-- The coercion from an elemental algebra to the full algebra as a `ClosedEmbedding`. -/ -theorem closedEmbedding_coe (x : A) : ClosedEmbedding ((↑) : elementalStarAlgebra R x → A) := +/-- The coercion from an elemental algebra to the full algebra as a `IsClosedEmbedding`. -/ +theorem isClosedEmbedding_coe (x : A) : IsClosedEmbedding ((↑) : elementalStarAlgebra R x → A) := { induced := rfl inj := Subtype.coe_injective isClosed_range := by @@ -217,6 +220,9 @@ theorem closedEmbedding_coe (x : A) : ClosedEmbedding ((↑) : elementalStarAlge rintro ⟨y, rfl⟩ exact y.prop, fun hy => ⟨⟨y, hy⟩, rfl⟩⟩ } +@[deprecated (since := "2024-10-20")] +alias closedEmbedding_coe := isClosedEmbedding_coe + @[elab_as_elim] theorem induction_on {x y : A} (hy : y ∈ elementalStarAlgebra R x) {P : (u : A) → u ∈ elementalStarAlgebra R x → Prop} diff --git a/Mathlib/Topology/Category/LightProfinite/Sequence.lean b/Mathlib/Topology/Category/LightProfinite/Sequence.lean index ae729f14b1979..adba9b5392add 100644 --- a/Mathlib/Topology/Category/LightProfinite/Sequence.lean +++ b/Mathlib/Topology/Category/LightProfinite/Sequence.lean @@ -29,8 +29,8 @@ noncomputable def natUnionInftyEmbedding : C(OnePoint ℕ, ℝ) where The continuous map from `ℕ∪{∞}` to `ℝ` sending `n` to `1/(n+1)` and `∞` to `0` is a closed embedding. -/ -lemma closedEmbedding_natUnionInftyEmbedding : ClosedEmbedding natUnionInftyEmbedding := by - refine closedEmbedding_of_continuous_injective_closed +lemma isClosedEmbedding_natUnionInftyEmbedding : IsClosedEmbedding natUnionInftyEmbedding := by + refine .of_continuous_injective_isClosedMap natUnionInftyEmbedding.continuous ?_ ?_ · rintro (_|n) (_|m) h · rfl @@ -45,7 +45,10 @@ lemma closedEmbedding_natUnionInftyEmbedding : ClosedEmbedding natUnionInftyEmbe rw [h] · exact fun _ hC => (hC.isCompact.image natUnionInftyEmbedding.continuous).isClosed -instance : MetrizableSpace (OnePoint ℕ) := closedEmbedding_natUnionInftyEmbedding.metrizableSpace +@[deprecated (since := "2024-10-20")] +alias closedEmbedding_natUnionInftyEmbedding := isClosedEmbedding_natUnionInftyEmbedding + +instance : MetrizableSpace (OnePoint ℕ) := isClosedEmbedding_natUnionInftyEmbedding.metrizableSpace /-- The one point compactification of the natural numbers as a light profinite set. -/ abbrev NatUnionInfty : LightProfinite := of (OnePoint ℕ) diff --git a/Mathlib/Topology/Category/Profinite/Nobeling.lean b/Mathlib/Topology/Category/Profinite/Nobeling.lean index 3807d5656a53a..275dd04efa2cf 100644 --- a/Mathlib/Topology/Category/Profinite/Nobeling.lean +++ b/Mathlib/Topology/Category/Profinite/Nobeling.lean @@ -1779,7 +1779,7 @@ def GoodProducts.Basis (hC : IsClosed C) : end Induction -variable {S : Profinite} {ι : S → I → Bool} (hι : ClosedEmbedding ι) +variable {S : Profinite} {ι : S → I → Bool} (hι : IsClosedEmbedding ι) include hι /-- @@ -1801,8 +1801,8 @@ def Nobeling.ι : S → ({C : Set S // IsClopen C} → Bool) := fun s C => decid open scoped Classical in /-- The map `Nobeling.ι` is a closed embedding. -/ -theorem Nobeling.embedding : ClosedEmbedding (Nobeling.ι S) := by - apply Continuous.closedEmbedding +theorem Nobeling.embedding : IsClosedEmbedding (Nobeling.ι S) := by + apply Continuous.isClosedEmbedding · dsimp (config := { unfoldPartialApp := true }) [ι] refine continuous_pi ?_ intro C diff --git a/Mathlib/Topology/Category/Stonean/Limits.lean b/Mathlib/Topology/Category/Stonean/Limits.lean index ed7151753cb73..2b9c500e98a2e 100644 --- a/Mathlib/Topology/Category/Stonean/Limits.lean +++ b/Mathlib/Topology/Category/Stonean/Limits.lean @@ -33,7 +33,7 @@ lemma extremallyDisconnected_preimage : ExtremallyDisconnected (i ⁻¹' (Set.ra ⟨IsClosed.preimage i.continuous (isCompact_range f.continuous).isClosed, IsOpen.preimage i.continuous hi.isOpen_range⟩ rw [← (closure U).preimage_image_eq Subtype.coe_injective, - ← h.1.closedEmbedding_subtype_val.closure_image_eq U] + ← h.1.isClosedEmbedding_subtypeVal.closure_image_eq U] exact isOpen_induced (ExtremallyDisconnected.open_closure _ (h.2.isOpenEmbedding_subtypeVal.isOpenMap U hU)) diff --git a/Mathlib/Topology/Clopen.lean b/Mathlib/Topology/Clopen.lean index 511935a8e3b98..8d4aa07d73434 100644 --- a/Mathlib/Topology/Clopen.lean +++ b/Mathlib/Topology/Clopen.lean @@ -118,7 +118,7 @@ theorem isClopen_range_inr : IsClopen (range (Sum.inr : Y → X ⊕ Y)) := theorem isClopen_range_sigmaMk {X : ι → Type*} [∀ i, TopologicalSpace (X i)] {i : ι} : IsClopen (Set.range (@Sigma.mk ι X i)) := - ⟨closedEmbedding_sigmaMk.isClosed_range, isOpenEmbedding_sigmaMk.isOpen_range⟩ + ⟨isClosedEmbedding_sigmaMk.isClosed_range, isOpenEmbedding_sigmaMk.isOpen_range⟩ protected theorem QuotientMap.isClopen_preimage {f : X → Y} (hf : QuotientMap f) {s : Set Y} : IsClopen (f ⁻¹' s) ↔ IsClopen s := diff --git a/Mathlib/Topology/Compactness/Compact.lean b/Mathlib/Topology/Compactness/Compact.lean index 0eb3000c7fbcf..ac4cbe1c1f71a 100644 --- a/Mathlib/Topology/Compactness/Compact.lean +++ b/Mathlib/Topology/Compactness/Compact.lean @@ -927,17 +927,23 @@ lemma Inducing.isCompact_preimage' {f : X → Y} (hf : Inducing f) {K : Set Y} (hf.isCompact_preimage_iff Kf).2 hK /-- The preimage of a compact set under a closed embedding is a compact set. -/ -theorem ClosedEmbedding.isCompact_preimage {f : X → Y} (hf : ClosedEmbedding f) +theorem IsClosedEmbedding.isCompact_preimage {f : X → Y} (hf : IsClosedEmbedding f) {K : Set Y} (hK : IsCompact K) : IsCompact (f ⁻¹' K) := hf.toInducing.isCompact_preimage (hf.isClosed_range) hK +@[deprecated (since := "2024-10-20")] +alias ClosedEmbedding.isCompact_preimage := IsClosedEmbedding.isCompact_preimage + /-- A closed embedding is proper, ie, inverse images of compact sets are contained in compacts. -Moreover, the preimage of a compact set is compact, see `ClosedEmbedding.isCompact_preimage`. -/ -theorem ClosedEmbedding.tendsto_cocompact {f : X → Y} (hf : ClosedEmbedding f) : +Moreover, the preimage of a compact set is compact, see `IsClosedEmbedding.isCompact_preimage`. -/ +theorem IsClosedEmbedding.tendsto_cocompact {f : X → Y} (hf : IsClosedEmbedding f) : Tendsto f (Filter.cocompact X) (Filter.cocompact Y) := Filter.hasBasis_cocompact.tendsto_right_iff.mpr fun _K hK => (hf.isCompact_preimage hK).compl_mem_cocompact +@[deprecated (since := "2024-10-20")] +alias ClosedEmbedding.tendsto_cocompact := IsClosedEmbedding.tendsto_cocompact + /-- Sets of subtype are compact iff the image under a coercion is. -/ theorem Subtype.isCompact_iff {p : X → Prop} {s : Set { x // p x }} : IsCompact s ↔ IsCompact ((↑) '' s : Set X) := @@ -956,14 +962,20 @@ theorem exists_nhds_ne_inf_principal_neBot (hs : IsCompact s) (hs' : s.Infinite) ∃ z ∈ s, (𝓝[≠] z ⊓ 𝓟 s).NeBot := hs'.exists_accPt_of_subset_isCompact hs Subset.rfl -protected theorem ClosedEmbedding.noncompactSpace [NoncompactSpace X] {f : X → Y} - (hf : ClosedEmbedding f) : NoncompactSpace Y := +protected theorem IsClosedEmbedding.noncompactSpace [NoncompactSpace X] {f : X → Y} + (hf : IsClosedEmbedding f) : NoncompactSpace Y := noncompactSpace_of_neBot hf.tendsto_cocompact.neBot -protected theorem ClosedEmbedding.compactSpace [h : CompactSpace Y] {f : X → Y} - (hf : ClosedEmbedding f) : CompactSpace X := +@[deprecated (since := "2024-10-20")] +alias ClosedEmbedding.noncompactSpace := IsClosedEmbedding.noncompactSpace + +protected theorem IsClosedEmbedding.compactSpace [h : CompactSpace Y] {f : X → Y} + (hf : IsClosedEmbedding f) : CompactSpace X := ⟨by rw [hf.toInducing.isCompact_iff, image_univ]; exact hf.isClosed_range.isCompact⟩ +@[deprecated (since := "2024-10-20")] +alias ClosedEmbedding.compactSpace := IsClosedEmbedding.compactSpace + theorem IsCompact.prod {t : Set Y} (hs : IsCompact s) (ht : IsCompact t) : IsCompact (s ×ˢ t) := by rw [isCompact_iff_ultrafilter_le_nhds'] at hs ht ⊢ @@ -981,7 +993,7 @@ instance (priority := 100) Finite.compactSpace [Finite X] : CompactSpace X where isCompact_univ := finite_univ.isCompact instance ULift.compactSpace [CompactSpace X] : CompactSpace (ULift.{v} X) := - ULift.closedEmbedding_down.compactSpace + ULift.isClosedEmbedding_down.compactSpace /-- The product of two compact spaces is compact. -/ instance [CompactSpace X] [CompactSpace Y] : CompactSpace (X × Y) := diff --git a/Mathlib/Topology/Compactness/Lindelof.lean b/Mathlib/Topology/Compactness/Lindelof.lean index 64c83c6ca1c47..9a5383d41086c 100644 --- a/Mathlib/Topology/Compactness/Lindelof.lean +++ b/Mathlib/Topology/Compactness/Lindelof.lean @@ -620,17 +620,24 @@ theorem Inducing.isLindelof_preimage {f : X → Y} (hf : Inducing f) (hf' : IsCl rwa [hf.isLindelof_iff, image_preimage_eq_inter_range] /-- The preimage of a Lindelöf set under a closed embedding is a Lindelöf set. -/ -theorem ClosedEmbedding.isLindelof_preimage {f : X → Y} (hf : ClosedEmbedding f) +theorem IsClosedEmbedding.isLindelof_preimage {f : X → Y} (hf : IsClosedEmbedding f) {K : Set Y} (hK : IsLindelof K) : IsLindelof (f ⁻¹' K) := hf.toInducing.isLindelof_preimage (hf.isClosed_range) hK +@[deprecated (since := "2024-10-20")] +alias ClosedEmbedding.isLindelof_preimage := IsClosedEmbedding.isLindelof_preimage + /-- A closed embedding is proper, ie, inverse images of Lindelöf sets are contained in Lindelöf. -Moreover, the preimage of a Lindelöf set is Lindelöf, see `ClosedEmbedding.isLindelof_preimage`. -/ -theorem ClosedEmbedding.tendsto_coLindelof {f : X → Y} (hf : ClosedEmbedding f) : +Moreover, the preimage of a Lindelöf set is Lindelöf, see +`IsClosedEmbedding.isLindelof_preimage`. -/ +theorem IsClosedEmbedding.tendsto_coLindelof {f : X → Y} (hf : IsClosedEmbedding f) : Tendsto f (Filter.coLindelof X) (Filter.coLindelof Y) := hasBasis_coLindelof.tendsto_right_iff.mpr fun _K hK => (hf.isLindelof_preimage hK).compl_mem_coLindelof +@[deprecated (since := "2024-10-20")] +alias ClosedEmbedding.tendsto_coLindelof := IsClosedEmbedding.tendsto_coLindelof + /-- Sets of subtype are Lindelöf iff the image under a coercion is. -/ theorem Subtype.isLindelof_iff {p : X → Prop} {s : Set { x // p x }} : IsLindelof s ↔ IsLindelof ((↑) '' s : Set X) := @@ -648,14 +655,20 @@ theorem IsLindelof.countable (hs : IsLindelof s) (hs' : DiscreteTopology s) : s. countable_coe_iff.mp (@countable_of_Lindelof_of_discrete _ _ (isLindelof_iff_LindelofSpace.mp hs) hs') -protected theorem ClosedEmbedding.nonLindelofSpace [NonLindelofSpace X] {f : X → Y} - (hf : ClosedEmbedding f) : NonLindelofSpace Y := +protected theorem IsClosedEmbedding.nonLindelofSpace [NonLindelofSpace X] {f : X → Y} + (hf : IsClosedEmbedding f) : NonLindelofSpace Y := nonLindelofSpace_of_neBot hf.tendsto_coLindelof.neBot -protected theorem ClosedEmbedding.LindelofSpace [h : LindelofSpace Y] {f : X → Y} - (hf : ClosedEmbedding f) : LindelofSpace X := +@[deprecated (since := "2024-10-20")] +alias ClosedEmbedding.nonLindelofSpace := IsClosedEmbedding.nonLindelofSpace + +protected theorem IsClosedEmbedding.LindelofSpace [h : LindelofSpace Y] {f : X → Y} + (hf : IsClosedEmbedding f) : LindelofSpace X := ⟨by rw [hf.toInducing.isLindelof_iff, image_univ]; exact hf.isClosed_range.isLindelof⟩ +@[deprecated (since := "2024-10-20")] +alias ClosedEmbedding.LindelofSpace := IsClosedEmbedding.LindelofSpace + /-- Countable topological spaces are Lindelof. -/ instance (priority := 100) Countable.LindelofSpace [Countable X] : LindelofSpace X where isLindelof_univ := countable_univ.isLindelof diff --git a/Mathlib/Topology/Compactness/LocallyCompact.lean b/Mathlib/Topology/Compactness/LocallyCompact.lean index 5bba2a34cb127..f479bd3f6e859 100644 --- a/Mathlib/Topology/Compactness/LocallyCompact.lean +++ b/Mathlib/Topology/Compactness/LocallyCompact.lean @@ -35,15 +35,18 @@ instance {ι : Type*} [Finite ι] {X : ι → Type*} [(i : ι) → TopologicalSp instance (priority := 100) [CompactSpace X] : WeaklyLocallyCompactSpace X where exists_compact_mem_nhds _ := ⟨univ, isCompact_univ, univ_mem⟩ -protected theorem ClosedEmbedding.weaklyLocallyCompactSpace [WeaklyLocallyCompactSpace Y] - {f : X → Y} (hf : ClosedEmbedding f) : WeaklyLocallyCompactSpace X where +protected theorem IsClosedEmbedding.weaklyLocallyCompactSpace [WeaklyLocallyCompactSpace Y] + {f : X → Y} (hf : IsClosedEmbedding f) : WeaklyLocallyCompactSpace X where exists_compact_mem_nhds x := let ⟨K, hK, hKx⟩ := exists_compact_mem_nhds (f x) ⟨f ⁻¹' K, hf.isCompact_preimage hK, hf.continuous.continuousAt hKx⟩ +@[deprecated (since := "2024-10-20")] +alias ClosedEmbedding.weaklyLocallyCompactSpace := IsClosedEmbedding.weaklyLocallyCompactSpace + protected theorem IsClosed.weaklyLocallyCompactSpace [WeaklyLocallyCompactSpace X] {s : Set X} (hs : IsClosed s) : WeaklyLocallyCompactSpace s := - (closedEmbedding_subtype_val hs).weaklyLocallyCompactSpace + hs.isClosedEmbedding_subtypeVal.weaklyLocallyCompactSpace theorem IsOpenQuotientMap.weaklyLocallyCompactSpace [WeaklyLocallyCompactSpace X] {f : X → Y} (hf : IsOpenQuotientMap f) : WeaklyLocallyCompactSpace Y where @@ -195,10 +198,13 @@ theorem Inducing.locallyCompactSpace [LocallyCompactSpace Y] {f : X → Y} (hf : rw [hf.isCompact_preimage_iff] exacts [hs.inter_right hZ, hUZ ▸ by gcongr] -protected theorem ClosedEmbedding.locallyCompactSpace [LocallyCompactSpace Y] {f : X → Y} - (hf : ClosedEmbedding f) : LocallyCompactSpace X := +protected theorem IsClosedEmbedding.locallyCompactSpace [LocallyCompactSpace Y] {f : X → Y} + (hf : IsClosedEmbedding f) : LocallyCompactSpace X := hf.toInducing.locallyCompactSpace hf.isClosed_range.isLocallyClosed +@[deprecated (since := "2024-10-20")] +alias ClosedEmbedding.locallyCompactSpace := IsClosedEmbedding.locallyCompactSpace + protected theorem IsOpenEmbedding.locallyCompactSpace [LocallyCompactSpace Y] {f : X → Y} (hf : IsOpenEmbedding f) : LocallyCompactSpace X := hf.toInducing.locallyCompactSpace hf.isOpen_range.isLocallyClosed diff --git a/Mathlib/Topology/Compactness/Paracompact.lean b/Mathlib/Topology/Compactness/Paracompact.lean index ea4ab6733b402..638f75142b792 100644 --- a/Mathlib/Topology/Compactness/Paracompact.lean +++ b/Mathlib/Topology/Compactness/Paracompact.lean @@ -107,8 +107,8 @@ theorem precise_refinement_set [ParacompactSpace X] {s : Set X} (hs : IsClosed s · simp only [iUnion_option, ← compl_subset_iff_union] at vc exact Subset.trans (subset_compl_comm.1 <| vu Option.none) vc -theorem ClosedEmbedding.paracompactSpace [ParacompactSpace Y] {e : X → Y} (he : ClosedEmbedding e) : - ParacompactSpace X where +theorem IsClosedEmbedding.paracompactSpace [ParacompactSpace Y] {e : X → Y} + (he : IsClosedEmbedding e) : ParacompactSpace X where locallyFinite_refinement α s ho hu := by choose U hUo hU using fun a ↦ he.isOpen_iff.1 (ho a) simp only [← hU] at hu ⊢ @@ -119,8 +119,11 @@ theorem ClosedEmbedding.paracompactSpace [ParacompactSpace Y] {e : X → Y} (he hVf.preimage_continuous he.continuous, fun a ↦ ⟨a, preimage_mono (hVU a)⟩⟩ simpa only [range_subset_iff, mem_iUnion, iUnion_eq_univ_iff] using heV +@[deprecated (since := "2024-10-20")] +alias ClosedEmbedding.paracompactSpace := IsClosedEmbedding.paracompactSpace + theorem Homeomorph.paracompactSpace_iff (e : X ≃ₜ Y) : ParacompactSpace X ↔ ParacompactSpace Y := - ⟨fun _ ↦ e.symm.closedEmbedding.paracompactSpace, fun _ ↦ e.closedEmbedding.paracompactSpace⟩ + ⟨fun _ ↦ e.symm.isClosedEmbedding.paracompactSpace, fun _ ↦ e.isClosedEmbedding.paracompactSpace⟩ /-- The product of a compact space and a paracompact space is a paracompact space. The formalization is based on https://dantopology.wordpress.com/2009/10/24/compact-x-paracompact-is-paracompact/ diff --git a/Mathlib/Topology/Compactness/SigmaCompact.lean b/Mathlib/Topology/Compactness/SigmaCompact.lean index ec5a6b0e2279f..9718f4ea99c27 100644 --- a/Mathlib/Topology/Compactness/SigmaCompact.lean +++ b/Mathlib/Topology/Compactness/SigmaCompact.lean @@ -249,17 +249,20 @@ instance [Countable ι] {X : ι → Type*} [∀ i, TopologicalSpace (X i)] refine ⟨max k n, k, le_max_left _ _, mem_image_of_mem _ ?_⟩ exact compactCovering_subset _ (le_max_right _ _) hn -protected theorem ClosedEmbedding.sigmaCompactSpace {e : Y → X} (he : ClosedEmbedding e) : +protected theorem IsClosedEmbedding.sigmaCompactSpace {e : Y → X} (he : IsClosedEmbedding e) : SigmaCompactSpace Y := ⟨⟨fun n => e ⁻¹' compactCovering X n, fun _ => he.isCompact_preimage (isCompact_compactCovering _ _), by rw [← preimage_iUnion, iUnion_compactCovering, preimage_univ]⟩⟩ +@[deprecated (since := "2024-10-20")] +alias ClosedEmbedding.sigmaCompactSpace := IsClosedEmbedding.sigmaCompactSpace + theorem IsClosed.sigmaCompactSpace {s : Set X} (hs : IsClosed s) : SigmaCompactSpace s := - (closedEmbedding_subtype_val hs).sigmaCompactSpace + hs.isClosedEmbedding_subtypeVal.sigmaCompactSpace instance [SigmaCompactSpace Y] : SigmaCompactSpace (ULift.{u} Y) := - ULift.closedEmbedding_down.sigmaCompactSpace + ULift.isClosedEmbedding_down.sigmaCompactSpace /-- If `X` is a `σ`-compact space, then a locally finite family of nonempty sets of `X` can have only countably many elements, `Set.Countable` version. -/ diff --git a/Mathlib/Topology/Constructions.lean b/Mathlib/Topology/Constructions.lean index 85297dd644e16..4a48f65c105f5 100644 --- a/Mathlib/Topology/Constructions.lean +++ b/Mathlib/Topology/Constructions.lean @@ -937,12 +937,18 @@ theorem isClosed_range_inr : IsClosed (range (inr : Y → X ⊕ Y)) := by rw [← isOpen_compl_iff, compl_range_inr] exact isOpen_range_inl -theorem closedEmbedding_inl : ClosedEmbedding (inl : X → X ⊕ Y) := +theorem isClosedEmbedding_inl : IsClosedEmbedding (inl : X → X ⊕ Y) := ⟨embedding_inl, isClosed_range_inl⟩ -theorem closedEmbedding_inr : ClosedEmbedding (inr : Y → X ⊕ Y) := +@[deprecated (since := "2024-10-20")] +alias closedEmbedding_inl := isClosedEmbedding_inl + +theorem isClosedEmbedding_inr : IsClosedEmbedding (inr : Y → X ⊕ Y) := ⟨embedding_inr, isClosed_range_inr⟩ +@[deprecated (since := "2024-10-20")] +alias closedEmbedding_inr := isClosedEmbedding_inr + theorem nhds_inl (x : X) : 𝓝 (inl x : X ⊕ Y) = map inl (𝓝 x) := (isOpenEmbedding_inl.map_nhds_eq _).symm @@ -981,7 +987,7 @@ theorem isClosedMap_sum {f : X ⊕ Y → Z} : IsClosedMap f ↔ (IsClosedMap fun a => f (.inl a)) ∧ IsClosedMap fun b => f (.inr b) := by constructor · intro h - exact ⟨h.comp closedEmbedding_inl.isClosedMap, h.comp closedEmbedding_inr.isClosedMap⟩ + exact ⟨h.comp isClosedEmbedding_inl.isClosedMap, h.comp isClosedEmbedding_inr.isClosedMap⟩ · rintro h Z hZ rw [isClosed_sum_iff] at hZ convert (h.1 _ hZ.1).union (h.2 _ hZ.2) @@ -1003,10 +1009,13 @@ theorem Inducing.of_codRestrict {f : X → Y} {t : Set Y} (ht : ∀ x, f x ∈ t theorem embedding_subtype_val : Embedding ((↑) : Subtype p → X) := ⟨inducing_subtype_val, Subtype.coe_injective⟩ -theorem closedEmbedding_subtype_val (h : IsClosed { a | p a }) : - ClosedEmbedding ((↑) : Subtype p → X) := +theorem IsClosedEmbedding.subtypeVal (h : IsClosed {a | p a}) : + IsClosedEmbedding ((↑) : Subtype p → X) := ⟨embedding_subtype_val, by rwa [Subtype.range_coe_subtype]⟩ +@[deprecated (since := "2024-10-20")] +alias closedEmbedding_subtype_val := IsClosedEmbedding.subtypeVal + @[continuity, fun_prop] theorem continuous_subtype_val : Continuous (@Subtype.val X p) := continuous_induced_dom @@ -1029,13 +1038,15 @@ theorem IsOpenMap.restrict {f : X → Y} (hf : IsOpenMap f) {s : Set X} (hs : Is IsOpenMap (s.restrict f) := hf.comp hs.isOpenMap_subtype_val -nonrec theorem IsClosed.closedEmbedding_subtype_val {s : Set X} (hs : IsClosed s) : - ClosedEmbedding ((↑) : s → X) := - closedEmbedding_subtype_val hs +lemma IsClosed.isClosedEmbedding_subtypeVal {s : Set X} (hs : IsClosed s) : + IsClosedEmbedding ((↑) : s → X) := .subtypeVal hs + +@[deprecated (since := "2024-10-20")] +alias IsClosed.closedEmbedding_subtype_val := IsClosed.isClosedEmbedding_subtypeVal theorem IsClosed.isClosedMap_subtype_val {s : Set X} (hs : IsClosed s) : IsClosedMap ((↑) : s → X) := - hs.closedEmbedding_subtype_val.isClosedMap + hs.isClosedEmbedding_subtypeVal.isClosedMap @[continuity, fun_prop] theorem Continuous.subtype_mk {f : Y → X} (h : Continuous f) (hp : ∀ x, p (f x)) : @@ -1547,12 +1558,15 @@ theorem isOpenEmbedding_sigmaMk {i : ι} : IsOpenEmbedding (@Sigma.mk ι σ i) : @[deprecated (since := "2024-10-18")] alias openEmbedding_sigmaMk := isOpenEmbedding_sigmaMk -theorem closedEmbedding_sigmaMk {i : ι} : ClosedEmbedding (@Sigma.mk ι σ i) := - closedEmbedding_of_continuous_injective_closed continuous_sigmaMk sigma_mk_injective +theorem isClosedEmbedding_sigmaMk {i : ι} : IsClosedEmbedding (@Sigma.mk ι σ i) := + .of_continuous_injective_isClosedMap continuous_sigmaMk sigma_mk_injective isClosedMap_sigmaMk +@[deprecated (since := "2024-10-20")] +alias closedEmbedding_sigmaMk := isClosedEmbedding_sigmaMk + theorem embedding_sigmaMk {i : ι} : Embedding (@Sigma.mk ι σ i) := - closedEmbedding_sigmaMk.1 + isClosedEmbedding_sigmaMk.1 theorem Sigma.nhds_mk (i : ι) (x : σ i) : 𝓝 (⟨i, x⟩ : Sigma σ) = Filter.map (Sigma.mk i) (𝓝 x) := (isOpenEmbedding_sigmaMk.map_nhds_eq x).symm @@ -1657,10 +1671,13 @@ theorem continuous_uLift_up [TopologicalSpace X] : Continuous (ULift.up : X → theorem embedding_uLift_down [TopologicalSpace X] : Embedding (ULift.down : ULift.{v, u} X → X) := ⟨⟨rfl⟩, ULift.down_injective⟩ -theorem ULift.closedEmbedding_down [TopologicalSpace X] : - ClosedEmbedding (ULift.down : ULift.{v, u} X → X) := +theorem ULift.isClosedEmbedding_down [TopologicalSpace X] : + IsClosedEmbedding (ULift.down : ULift.{v, u} X → X) := ⟨embedding_uLift_down, by simp only [ULift.down_surjective.range_eq, isClosed_univ]⟩ +@[deprecated (since := "2024-10-20")] +alias ULift.closedEmbedding_down := ULift.isClosedEmbedding_down + instance [TopologicalSpace X] [DiscreteTopology X] : DiscreteTopology (ULift X) := embedding_uLift_down.discreteTopology diff --git a/Mathlib/Topology/ContinuousMap/ContinuousMapZero.lean b/Mathlib/Topology/ContinuousMap/ContinuousMapZero.lean index 6d699f5ba3042..6bb261df3d76d 100644 --- a/Mathlib/Topology/ContinuousMap/ContinuousMapZero.lean +++ b/Mathlib/Topology/ContinuousMap/ContinuousMapZero.lean @@ -98,13 +98,16 @@ instance instContinuousEvalConst : ContinuousEvalConst C(X, R)₀ X R := instance instContinuousEval [LocallyCompactPair X R] : ContinuousEval C(X, R)₀ X R := .of_continuous_forget embedding_toContinuousMap.continuous -lemma closedEmbedding_toContinuousMap [T1Space R] : - ClosedEmbedding ((↑) : C(X, R)₀ → C(X, R)) where +lemma isClosedEmbedding_toContinuousMap [T1Space R] : + IsClosedEmbedding ((↑) : C(X, R)₀ → C(X, R)) where toEmbedding := embedding_toContinuousMap isClosed_range := by rw [range_toContinuousMap] exact isClosed_singleton.preimage <| continuous_eval_const 0 +@[deprecated (since := "2024-10-20")] +alias closedEmbedding_toContinuousMap := isClosedEmbedding_toContinuousMap + @[fun_prop] lemma continuous_comp_left {X Y Z : Type*} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] [Zero X] [Zero Y] [Zero Z] (f : C(X, Y)₀) : @@ -288,7 +291,7 @@ alias uniformEmbedding_toContinuousMap := isUniformEmbedding_toContinuousMap instance [T1Space R] [CompleteSpace C(X, R)] : CompleteSpace C(X, R)₀ := completeSpace_iff_isComplete_range isUniformEmbedding_toContinuousMap.isUniformInducing - |>.mpr closedEmbedding_toContinuousMap.isClosed_range.isComplete + |>.mpr isClosedEmbedding_toContinuousMap.isClosed_range.isComplete lemma isUniformEmbedding_comp {Y : Type*} [UniformSpace Y] [Zero Y] (g : C(Y, R)₀) (hg : IsUniformEmbedding g) : IsUniformEmbedding (g.comp · : C(X, Y)₀ → C(X, R)₀) := diff --git a/Mathlib/Topology/ContinuousMap/StoneWeierstrass.lean b/Mathlib/Topology/ContinuousMap/StoneWeierstrass.lean index ecfa6bfa4e8f4..7d274f86f8aab 100644 --- a/Mathlib/Topology/ContinuousMap/StoneWeierstrass.lean +++ b/Mathlib/Topology/ContinuousMap/StoneWeierstrass.lean @@ -588,8 +588,8 @@ lemma ContinuousMapZero.adjoin_id_dense {s : Set 𝕜} [Zero s] (h0 : ((0 : s) : [CompactSpace s] : Dense (adjoin 𝕜 {(.id h0 : C(s, 𝕜)₀)} : Set C(s, 𝕜)₀) := by have h0' : 0 ∈ s := h0 ▸ (0 : s).property rw [dense_iff_closure_eq, - ← closedEmbedding_toContinuousMap.injective.preimage_image (closure _), - ← closedEmbedding_toContinuousMap.closure_image_eq, ← coe_toContinuousMapHom, + ← isClosedEmbedding_toContinuousMap.injective.preimage_image (closure _), + ← isClosedEmbedding_toContinuousMap.closure_image_eq, ← coe_toContinuousMapHom, ← NonUnitalStarSubalgebra.coe_map, NonUnitalStarAlgHom.map_adjoin_singleton, toContinuousMapHom_apply, toContinuousMap_id h0, ← ContinuousMap.ker_evalStarAlgHom_eq_closure_adjoin_id s h0'] diff --git a/Mathlib/Topology/Defs/Induced.lean b/Mathlib/Topology/Defs/Induced.lean index 0aaa1d46b8de9..0e1df6198a3bd 100644 --- a/Mathlib/Topology/Defs/Induced.lean +++ b/Mathlib/Topology/Defs/Induced.lean @@ -32,7 +32,7 @@ as well as topology inducing maps, topological embeddings, and quotient maps. if it is an embedding and its range is open. An open embedding is an open map. -* `ClosedEmbedding`: a map `f : X → Y` is an *open embedding*, +* `IsClosedEmbedding`: a map `f : X → Y` is an *open embedding*, if it is an embedding and its range is open. An open embedding is an open map. @@ -121,10 +121,13 @@ alias OpenEmbedding := IsOpenEmbedding /-- A closed embedding is an embedding with closed image. -/ @[mk_iff] -structure ClosedEmbedding (f : X → Y) extends Embedding f : Prop where +structure IsClosedEmbedding (f : X → Y) extends Embedding f : Prop where /-- The range of a closed embedding is a closed set. -/ isClosed_range : IsClosed <| range f +@[deprecated (since := "2024-10-20")] +alias ClosedEmbedding := IsClosedEmbedding + /-- A function between topological spaces is a quotient map if it is surjective, and for all `s : Set Y`, `s` is open iff its preimage is an open set. -/ @[mk_iff quotientMap_iff'] diff --git a/Mathlib/Topology/DiscreteSubset.lean b/Mathlib/Topology/DiscreteSubset.lean index 5238fbf4e8e1e..825509afa0b34 100644 --- a/Mathlib/Topology/DiscreteSubset.lean +++ b/Mathlib/Topology/DiscreteSubset.lean @@ -69,7 +69,7 @@ lemma tendsto_cofinite_cocompact_of_discrete [DiscreteTopology X] lemma IsClosed.tendsto_coe_cofinite_of_discreteTopology {s : Set X} (hs : IsClosed s) (_hs' : DiscreteTopology s) : Tendsto ((↑) : s → X) cofinite (cocompact _) := - tendsto_cofinite_cocompact_of_discrete hs.closedEmbedding_subtype_val.tendsto_cocompact + tendsto_cofinite_cocompact_of_discrete hs.isClosedEmbedding_subtypeVal.tendsto_cocompact lemma IsClosed.tendsto_coe_cofinite_iff [T1Space X] [WeaklyLocallyCompactSpace X] {s : Set X} (hs : IsClosed s) : diff --git a/Mathlib/Topology/FiberBundle/Basic.lean b/Mathlib/Topology/FiberBundle/Basic.lean index 39a1e8d2edc2b..3fb66e7803f27 100644 --- a/Mathlib/Topology/FiberBundle/Basic.lean +++ b/Mathlib/Topology/FiberBundle/Basic.lean @@ -261,12 +261,15 @@ theorem continuous_totalSpaceMk (x : B) : Continuous (@TotalSpace.mk B F E x) := theorem totalSpaceMk_embedding (x : B) : Embedding (@TotalSpace.mk B F E x) := ⟨totalSpaceMk_inducing F E x, TotalSpace.mk_injective x⟩ -theorem totalSpaceMk_closedEmbedding [T1Space B] (x : B) : - ClosedEmbedding (@TotalSpace.mk B F E x) := +theorem totalSpaceMk_isClosedEmbedding [T1Space B] (x : B) : + IsClosedEmbedding (@TotalSpace.mk B F E x) := ⟨totalSpaceMk_embedding F E x, by rw [TotalSpace.range_mk] exact isClosed_singleton.preimage <| continuous_proj F E⟩ +@[deprecated (since := "2024-10-20")] +alias totalSpaceMk_closedEmbedding := totalSpaceMk_isClosedEmbedding + variable {E F} @[simp, mfld_simps] diff --git a/Mathlib/Topology/Homeomorph.lean b/Mathlib/Topology/Homeomorph.lean index ce885a2ca891d..d1f3d9a2db354 100644 --- a/Mathlib/Topology/Homeomorph.lean +++ b/Mathlib/Topology/Homeomorph.lean @@ -338,14 +338,17 @@ theorem isOpenEmbedding (h : X ≃ₜ Y) : IsOpenEmbedding h := @[deprecated (since := "2024-10-18")] alias openEmbedding := isOpenEmbedding -protected theorem closedEmbedding (h : X ≃ₜ Y) : ClosedEmbedding h := - closedEmbedding_of_embedding_closed h.embedding h.isClosedMap +theorem isClosedEmbedding (h : X ≃ₜ Y) : IsClosedEmbedding h := + .of_embedding_closed h.embedding h.isClosedMap + +@[deprecated (since := "2024-10-20")] +alias closedEmbedding := isClosedEmbedding protected theorem normalSpace [NormalSpace X] (h : X ≃ₜ Y) : NormalSpace Y := - h.symm.closedEmbedding.normalSpace + h.symm.isClosedEmbedding.normalSpace protected theorem t4Space [T4Space X] (h : X ≃ₜ Y) : T4Space Y := - h.symm.closedEmbedding.t4Space + h.symm.isClosedEmbedding.t4Space theorem preimage_closure (h : X ≃ₜ Y) (s : Set Y) : h ⁻¹' closure s = closure (h ⁻¹' s) := h.isOpenMap.preimage_closure_eq_closure_preimage h.continuous _ @@ -368,7 +371,7 @@ theorem image_frontier (h : X ≃ₜ Y) (s : Set X) : h '' frontier s = frontier @[to_additive] theorem _root_.HasCompactMulSupport.comp_homeomorph {M} [One M] {f : Y → M} (hf : HasCompactMulSupport f) (φ : X ≃ₜ Y) : HasCompactMulSupport (f ∘ φ) := - hf.comp_closedEmbedding φ.closedEmbedding + hf.comp_isClosedEmbedding φ.isClosedEmbedding @[simp] theorem map_nhds_eq (h : X ≃ₜ Y) (x : X) : map h (𝓝 x) = 𝓝 (h x) := @@ -405,7 +408,7 @@ the domain is a locally compact space. -/ theorem locallyCompactSpace_iff (h : X ≃ₜ Y) : LocallyCompactSpace X ↔ LocallyCompactSpace Y := by exact ⟨fun _ => h.symm.isOpenEmbedding.locallyCompactSpace, - fun _ => h.closedEmbedding.locallyCompactSpace⟩ + fun _ => h.isClosedEmbedding.locallyCompactSpace⟩ /-- If a bijective map `e : X ≃ Y` is continuous and open, then it is a homeomorphism. -/ @[simps toEquiv] @@ -912,9 +915,10 @@ protected lemma inducing : Inducing f := (hf.homeomorph f).inducing protected lemma quotientMap : QuotientMap f := (hf.homeomorph f).quotientMap protected lemma embedding : Embedding f := (hf.homeomorph f).embedding lemma isOpenEmbedding : IsOpenEmbedding f := (hf.homeomorph f).isOpenEmbedding -protected lemma closedEmbedding : ClosedEmbedding f := (hf.homeomorph f).closedEmbedding +lemma isClosedEmbedding : IsClosedEmbedding f := (hf.homeomorph f).isClosedEmbedding lemma isDenseEmbedding : IsDenseEmbedding f := (hf.homeomorph f).isDenseEmbedding +@[deprecated (since := "2024-10-20")] alias closedEmbedding := isClosedEmbedding @[deprecated (since := "2024-10-18")] alias openEmbedding := isOpenEmbedding diff --git a/Mathlib/Topology/Instances/CantorSet.lean b/Mathlib/Topology/Instances/CantorSet.lean index 22167d7ac0fbb..69820c6161f0a 100644 --- a/Mathlib/Topology/Instances/CantorSet.lean +++ b/Mathlib/Topology/Instances/CantorSet.lean @@ -86,8 +86,8 @@ lemma isClosed_preCantorSet (n : ℕ) : IsClosed (preCantorSet n) := by | zero => exact isClosed_Icc | succ n ih => refine IsClosed.union ?_ ?_ - · simpa [f, div_eq_inv_mul] using f.closedEmbedding.closed_iff_image_closed.mp ih - · simpa [g, f, div_eq_inv_mul] using g.closedEmbedding.closed_iff_image_closed.mp ih + · simpa [f, div_eq_inv_mul] using f.isClosedEmbedding.closed_iff_image_closed.mp ih + · simpa [g, f, div_eq_inv_mul] using g.isClosedEmbedding.closed_iff_image_closed.mp ih /-- The ternary Cantor set is closed. -/ lemma isClosed_cantorSet : IsClosed cantorSet := diff --git a/Mathlib/Topology/Instances/EReal.lean b/Mathlib/Topology/Instances/EReal.lean index 1b66243074048..89c410a5d2b73 100644 --- a/Mathlib/Topology/Instances/EReal.lean +++ b/Mathlib/Topology/Instances/EReal.lean @@ -96,9 +96,12 @@ theorem embedding_coe_ennreal : Embedding ((↑) : ℝ≥0∞ → EReal) := coe_ennreal_strictMono.embedding_of_ordConnected <| by rw [range_coe_ennreal]; exact ordConnected_Ici -theorem closedEmbedding_coe_ennreal : ClosedEmbedding ((↑) : ℝ≥0∞ → EReal) := +theorem isClosedEmbedding_coe_ennreal : IsClosedEmbedding ((↑) : ℝ≥0∞ → EReal) := ⟨embedding_coe_ennreal, by rw [range_coe_ennreal]; exact isClosed_Ici⟩ +@[deprecated (since := "2024-10-20")] +alias closedEmbedding_coe_ennreal := isClosedEmbedding_coe_ennreal + @[norm_cast] theorem tendsto_coe_ennreal {α : Type*} {f : Filter α} {m : α → ℝ≥0∞} {a : ℝ≥0∞} : Tendsto (fun a => (m a : EReal)) f (𝓝 ↑a) ↔ Tendsto m f (𝓝 a) := diff --git a/Mathlib/Topology/Instances/Int.lean b/Mathlib/Topology/Instances/Int.lean index 5ee25d7b929f2..536b95e356e15 100644 --- a/Mathlib/Topology/Instances/Int.lean +++ b/Mathlib/Topology/Instances/Int.lean @@ -45,8 +45,11 @@ theorem isUniformEmbedding_coe_real : IsUniformEmbedding ((↑) : ℤ → ℝ) : @[deprecated (since := "2024-10-01")] alias uniformEmbedding_coe_real := isUniformEmbedding_coe_real -theorem closedEmbedding_coe_real : ClosedEmbedding ((↑) : ℤ → ℝ) := - closedEmbedding_of_pairwise_le_dist zero_lt_one pairwise_one_le_dist +theorem isClosedEmbedding_coe_real : IsClosedEmbedding ((↑) : ℤ → ℝ) := + isClosedEmbedding_of_pairwise_le_dist zero_lt_one pairwise_one_le_dist + +@[deprecated (since := "2024-10-20")] +alias closedEmbedding_coe_real := isClosedEmbedding_coe_real instance : MetricSpace ℤ := Int.isUniformEmbedding_coe_real.comapMetricSpace _ diff --git a/Mathlib/Topology/Instances/Irrational.lean b/Mathlib/Topology/Instances/Irrational.lean index fc7c1f606e691..7d4b39145e504 100644 --- a/Mathlib/Topology/Instances/Irrational.lean +++ b/Mathlib/Topology/Instances/Irrational.lean @@ -71,7 +71,7 @@ instance : DenselyOrdered { x // Irrational x } := theorem eventually_forall_le_dist_cast_div (hx : Irrational x) (n : ℕ) : ∀ᶠ ε : ℝ in 𝓝 0, ∀ m : ℤ, ε ≤ dist x (m / n) := by have A : IsClosed (range (fun m => (n : ℝ)⁻¹ * m : ℤ → ℝ)) := - ((isClosedMap_smul₀ (n⁻¹ : ℝ)).comp Int.closedEmbedding_coe_real.isClosedMap).isClosed_range + ((isClosedMap_smul₀ (n⁻¹ : ℝ)).comp Int.isClosedEmbedding_coe_real.isClosedMap).isClosed_range have B : x ∉ range (fun m => (n : ℝ)⁻¹ * m : ℤ → ℝ) := by rintro ⟨m, rfl⟩ simp at hx diff --git a/Mathlib/Topology/Instances/NNReal.lean b/Mathlib/Topology/Instances/NNReal.lean index 6f0a140493505..16383d3381a75 100644 --- a/Mathlib/Topology/Instances/NNReal.lean +++ b/Mathlib/Topology/Instances/NNReal.lean @@ -289,7 +289,7 @@ end Monotone instance instProperSpace : ProperSpace ℝ≥0 where isCompact_closedBall x r := by - have emb : ClosedEmbedding ((↑) : ℝ≥0 → ℝ) := Isometry.closedEmbedding fun _ ↦ congrFun rfl + have emb : IsClosedEmbedding ((↑) : ℝ≥0 → ℝ) := Isometry.isClosedEmbedding fun _ ↦ congrFun rfl exact emb.isCompact_preimage (K := Metric.closedBall x r) (isCompact_closedBall _ _) end NNReal diff --git a/Mathlib/Topology/Instances/Nat.lean b/Mathlib/Topology/Instances/Nat.lean index e90ce1cbe65b7..c5719cb578fc2 100644 --- a/Mathlib/Topology/Instances/Nat.lean +++ b/Mathlib/Topology/Instances/Nat.lean @@ -37,8 +37,11 @@ theorem isUniformEmbedding_coe_real : IsUniformEmbedding ((↑) : ℕ → ℝ) : @[deprecated (since := "2024-10-01")] alias uniformEmbedding_coe_real := isUniformEmbedding_coe_real -theorem closedEmbedding_coe_real : ClosedEmbedding ((↑) : ℕ → ℝ) := - closedEmbedding_of_pairwise_le_dist zero_lt_one pairwise_one_le_dist +theorem isClosedEmbedding_coe_real : IsClosedEmbedding ((↑) : ℕ → ℝ) := + isClosedEmbedding_of_pairwise_le_dist zero_lt_one pairwise_one_le_dist + +@[deprecated (since := "2024-10-20")] +alias closedEmbedding_coe_real := isClosedEmbedding_coe_real instance : MetricSpace ℕ := Nat.isUniformEmbedding_coe_real.comapMetricSpace _ diff --git a/Mathlib/Topology/Instances/Rat.lean b/Mathlib/Topology/Instances/Rat.lean index 715361b5de0d9..8847d81d5d088 100644 --- a/Mathlib/Topology/Instances/Rat.lean +++ b/Mathlib/Topology/Instances/Rat.lean @@ -60,8 +60,11 @@ theorem Nat.isUniformEmbedding_coe_rat : IsUniformEmbedding ((↑) : ℕ → ℚ @[deprecated (since := "2024-10-01")] alias Nat.uniformEmbedding_coe_rat := Nat.isUniformEmbedding_coe_rat -theorem Nat.closedEmbedding_coe_rat : ClosedEmbedding ((↑) : ℕ → ℚ) := - closedEmbedding_of_pairwise_le_dist zero_lt_one <| by simpa using Nat.pairwise_one_le_dist +theorem Nat.isClosedEmbedding_coe_rat : IsClosedEmbedding ((↑) : ℕ → ℚ) := + isClosedEmbedding_of_pairwise_le_dist zero_lt_one <| by simpa using Nat.pairwise_one_le_dist + +@[deprecated (since := "2024-10-20")] +alias Nat.closedEmbedding_coe_rat := Nat.isClosedEmbedding_coe_rat @[norm_cast, simp] theorem Int.dist_cast_rat (x y : ℤ) : dist (x : ℚ) y = dist x y := by @@ -73,12 +76,15 @@ theorem Int.isUniformEmbedding_coe_rat : IsUniformEmbedding ((↑) : ℤ → ℚ @[deprecated (since := "2024-10-01")] alias Int.uniformEmbedding_coe_rat := Int.isUniformEmbedding_coe_rat -theorem Int.closedEmbedding_coe_rat : ClosedEmbedding ((↑) : ℤ → ℚ) := - closedEmbedding_of_pairwise_le_dist zero_lt_one <| by simpa using Int.pairwise_one_le_dist +theorem Int.isClosedEmbedding_coe_rat : IsClosedEmbedding ((↑) : ℤ → ℚ) := + isClosedEmbedding_of_pairwise_le_dist zero_lt_one <| by simpa using Int.pairwise_one_le_dist + +@[deprecated (since := "2024-10-20")] +alias Int.closedEmbedding_coe_rat := Int.isClosedEmbedding_coe_rat namespace Rat -instance : NoncompactSpace ℚ := Int.closedEmbedding_coe_rat.noncompactSpace +instance : NoncompactSpace ℚ := Int.isClosedEmbedding_coe_rat.noncompactSpace theorem uniformContinuous_add : UniformContinuous fun p : ℚ × ℚ => p.1 + p.2 := Rat.isUniformEmbedding_coe_real.isUniformInducing.uniformContinuous_iff.2 <| by diff --git a/Mathlib/Topology/Instances/Real.lean b/Mathlib/Topology/Instances/Real.lean index 927f93d6bf83a..c41f014575333 100644 --- a/Mathlib/Topology/Instances/Real.lean +++ b/Mathlib/Topology/Instances/Real.lean @@ -28,7 +28,7 @@ universe u v w variable {α : Type u} {β : Type v} {γ : Type w} -instance : NoncompactSpace ℝ := Int.closedEmbedding_coe_real.noncompactSpace +instance : NoncompactSpace ℝ := Int.isClosedEmbedding_coe_real.noncompactSpace theorem Real.uniformContinuous_add : UniformContinuous fun p : ℝ × ℝ => p.1 + p.2 := Metric.uniformContinuous_iff.2 fun _ε ε0 => diff --git a/Mathlib/Topology/KrullDimension.lean b/Mathlib/Topology/KrullDimension.lean index b5eee2e982ae4..9a25326df9e83 100644 --- a/Mathlib/Topology/KrullDimension.lean +++ b/Mathlib/Topology/KrullDimension.lean @@ -40,7 +40,7 @@ def IrreducibleCloseds.map {f : X → Y} (hf1 : Continuous f) (hf2 : IsClosedMap /-- Taking images under a closed embedding is strictly monotone on the preorder of irreducible closeds. -/ -lemma IrreducibleCloseds.map_strictMono {f : X → Y} (hf : ClosedEmbedding f) : +lemma IrreducibleCloseds.map_strictMono {f : X → Y} (hf : IsClosedEmbedding f) : StrictMono (IrreducibleCloseds.map hf.continuous hf.isClosedMap) := fun ⦃_ _⦄ UltV ↦ hf.inj.image_strictMono UltV @@ -48,16 +48,19 @@ lemma IrreducibleCloseds.map_strictMono {f : X → Y} (hf : ClosedEmbedding f) : If `f : X → Y` is a closed embedding, then the Krull dimension of `X` is less than or equal to the Krull dimension of `Y`. -/ -theorem ClosedEmbedding.topologicalKrullDim_le (f : X → Y) (hf : ClosedEmbedding f) : +theorem IsClosedEmbedding.topologicalKrullDim_le (f : X → Y) (hf : IsClosedEmbedding f) : topologicalKrullDim X ≤ topologicalKrullDim Y := krullDim_le_of_strictMono _ (IrreducibleCloseds.map_strictMono hf) +@[deprecated (since := "2024-10-20")] +alias ClosedEmbedding.topologicalKrullDim_le := IsClosedEmbedding.topologicalKrullDim_le + /-- The topological Krull dimension is invariant under homeomorphisms -/ theorem IsHomeomorph.topologicalKrullDim_eq (f : X → Y) (h : IsHomeomorph f) : topologicalKrullDim X = topologicalKrullDim Y := have fwd : topologicalKrullDim X ≤ topologicalKrullDim Y := - ClosedEmbedding.topologicalKrullDim_le f h.closedEmbedding + IsClosedEmbedding.topologicalKrullDim_le f h.isClosedEmbedding have bwd : topologicalKrullDim Y ≤ topologicalKrullDim X := - ClosedEmbedding.topologicalKrullDim_le (h.homeomorph f).symm - (h.homeomorph f).symm.closedEmbedding + IsClosedEmbedding.topologicalKrullDim_le (h.homeomorph f).symm + (h.homeomorph f).symm.isClosedEmbedding le_antisymm fwd bwd diff --git a/Mathlib/Topology/LocalAtTarget.lean b/Mathlib/Topology/LocalAtTarget.lean index 4e1a6f06f340a..c1e9f1b8153f2 100644 --- a/Mathlib/Topology/LocalAtTarget.lean +++ b/Mathlib/Topology/LocalAtTarget.lean @@ -13,7 +13,7 @@ We show that the following properties of continuous maps are local at the target - `Inducing` - `Embedding` - `IsOpenEmbedding` -- `ClosedEmbedding` +- `IsClosedEmbedding` -/ @@ -53,12 +53,18 @@ alias IsOpenEmbedding.restrictPreimage := Set.restrictPreimage_isOpenEmbedding @[deprecated (since := "2024-10-18")] alias OpenEmbedding.restrictPreimage := IsOpenEmbedding.restrictPreimage -theorem Set.restrictPreimage_closedEmbedding (s : Set β) (h : ClosedEmbedding f) : - ClosedEmbedding (s.restrictPreimage f) := +theorem Set.restrictPreimage_isClosedEmbedding (s : Set β) (h : IsClosedEmbedding f) : + IsClosedEmbedding (s.restrictPreimage f) := ⟨h.1.restrictPreimage s, (s.range_restrictPreimage f).symm ▸ inducing_subtype_val.isClosed_preimage _ h.isClosed_range⟩ -alias ClosedEmbedding.restrictPreimage := Set.restrictPreimage_closedEmbedding +@[deprecated (since := "2024-10-20")] +alias Set.restrictPreimage_closedEmbedding := Set.restrictPreimage_isClosedEmbedding + +alias IsClosedEmbedding.restrictPreimage := Set.restrictPreimage_isClosedEmbedding + +@[deprecated (since := "2024-10-20")] +alias ClosedEmbedding.restrictPreimage := IsClosedEmbedding.restrictPreimage theorem IsClosedMap.restrictPreimage (H : IsClosedMap f) (s : Set β) : IsClosedMap (s.restrictPreimage f) := by @@ -184,11 +190,15 @@ theorem isOpenEmbedding_iff_isOpenEmbedding_of_iSup_eq_top (h : Continuous f) : alias openEmbedding_iff_openEmbedding_of_iSup_eq_top := isOpenEmbedding_iff_isOpenEmbedding_of_iSup_eq_top -theorem closedEmbedding_iff_closedEmbedding_of_iSup_eq_top (h : Continuous f) : - ClosedEmbedding f ↔ ∀ i, ClosedEmbedding ((U i).1.restrictPreimage f) := by - simp_rw [closedEmbedding_iff] +theorem isClosedEmbedding_iff_isClosedEmbedding_of_iSup_eq_top (h : Continuous f) : + IsClosedEmbedding f ↔ ∀ i, IsClosedEmbedding ((U i).1.restrictPreimage f) := by + simp_rw [isClosedEmbedding_iff] rw [forall_and] apply and_congr · apply embedding_iff_embedding_of_iSup_eq_top <;> assumption · simp_rw [Set.range_restrictPreimage] apply isClosed_iff_coe_preimage_of_iSup_eq_top hU + +@[deprecated (since := "2024-10-20")] +alias closedEmbedding_iff_closedEmbedding_of_iSup_eq_top := + isClosedEmbedding_iff_isClosedEmbedding_of_iSup_eq_top diff --git a/Mathlib/Topology/Maps/Basic.lean b/Mathlib/Topology/Maps/Basic.lean index c4aaea387693f..7933f9ee54473 100644 --- a/Mathlib/Topology/Maps/Basic.lean +++ b/Mathlib/Topology/Maps/Basic.lean @@ -22,7 +22,7 @@ This file introduces the following properties of a map `f : X → Y` between top a subspace of `Y`. * `IsOpenEmbedding f` means `f` is an embedding with open image, so it identifies `X` with an open subspace of `Y`. Equivalently, `f` is an embedding and an open map. -* `ClosedEmbedding f` similarly means `f` is an embedding with closed image, so it identifies +* `IsClosedEmbedding f` similarly means `f` is an embedding with closed image, so it identifies `X` with a closed subspace of `Y`. Equivalently, `f` is an embedding and a closed map. * `QuotientMap f` is the dual condition to `Embedding f`: `f` is surjective and the topology @@ -603,59 +603,69 @@ end IsOpenEmbedding end IsOpenEmbedding -section ClosedEmbedding +section IsClosedEmbedding variable [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] -namespace ClosedEmbedding +namespace IsClosedEmbedding -theorem tendsto_nhds_iff {g : ι → X} {l : Filter ι} {x : X} (hf : ClosedEmbedding f) : +theorem tendsto_nhds_iff {g : ι → X} {l : Filter ι} {x : X} (hf : IsClosedEmbedding f) : Tendsto g l (𝓝 x) ↔ Tendsto (f ∘ g) l (𝓝 (f x)) := hf.toEmbedding.tendsto_nhds_iff -theorem continuous (hf : ClosedEmbedding f) : Continuous f := +theorem continuous (hf : IsClosedEmbedding f) : Continuous f := hf.toEmbedding.continuous -theorem isClosedMap (hf : ClosedEmbedding f) : IsClosedMap f := +theorem isClosedMap (hf : IsClosedEmbedding f) : IsClosedMap f := hf.toEmbedding.toInducing.isClosedMap hf.isClosed_range -theorem closed_iff_image_closed (hf : ClosedEmbedding f) {s : Set X} : +theorem closed_iff_image_closed (hf : IsClosedEmbedding f) {s : Set X} : IsClosed s ↔ IsClosed (f '' s) := ⟨hf.isClosedMap s, fun h => by rw [← preimage_image_eq s hf.inj] exact h.preimage hf.continuous⟩ -theorem closed_iff_preimage_closed (hf : ClosedEmbedding f) {s : Set Y} +theorem closed_iff_preimage_closed (hf : IsClosedEmbedding f) {s : Set Y} (hs : s ⊆ range f) : IsClosed s ↔ IsClosed (f ⁻¹' s) := by rw [hf.closed_iff_image_closed, image_preimage_eq_of_subset hs] -theorem _root_.closedEmbedding_of_embedding_closed (h₁ : Embedding f) (h₂ : IsClosedMap f) : - ClosedEmbedding f := +theorem _root_.IsClosedEmbedding.of_embedding_closed (h₁ : Embedding f) (h₂ : IsClosedMap f) : + IsClosedEmbedding f := ⟨h₁, image_univ (f := f) ▸ h₂ univ isClosed_univ⟩ -theorem _root_.closedEmbedding_of_continuous_injective_closed (h₁ : Continuous f) (h₂ : Injective f) - (h₃ : IsClosedMap f) : ClosedEmbedding f := by - refine closedEmbedding_of_embedding_closed ⟨⟨?_⟩, h₂⟩ h₃ +@[deprecated (since := "2024-10-20")] +alias _root_.closedEmbedding_of_embedding_closed := _root_.IsClosedEmbedding.of_embedding_closed + +lemma _root_.IsClosedEmbedding.of_continuous_injective_isClosedMap (h₁ : Continuous f) + (h₂ : Injective f) (h₃ : IsClosedMap f) : IsClosedEmbedding f := by + refine .of_embedding_closed ⟨⟨?_⟩, h₂⟩ h₃ refine h₁.le_induced.antisymm fun s hs => ?_ refine ⟨(f '' sᶜ)ᶜ, (h₃ _ hs.isClosed_compl).isOpen_compl, ?_⟩ rw [preimage_compl, preimage_image_eq _ h₂, compl_compl] -theorem _root_.closedEmbedding_id : ClosedEmbedding (@id X) := +@[deprecated (since := "2024-10-20")] +alias _root_.closedEmbedding_of_continuous_injective_closed := + IsClosedEmbedding.of_continuous_injective_isClosedMap + +theorem _root_.isClosedEmbedding_id : IsClosedEmbedding (@id X) := ⟨embedding_id, IsClosedMap.id.isClosed_range⟩ -theorem comp (hg : ClosedEmbedding g) (hf : ClosedEmbedding f) : - ClosedEmbedding (g ∘ f) := +@[deprecated (since := "2024-10-20")] +alias _root_.closedEmbedding_id := _root_.isClosedEmbedding_id + +theorem comp (hg : IsClosedEmbedding g) (hf : IsClosedEmbedding f) : + IsClosedEmbedding (g ∘ f) := ⟨hg.toEmbedding.comp hf.toEmbedding, (hg.isClosedMap.comp hf.isClosedMap).isClosed_range⟩ -theorem of_comp_iff (hg : ClosedEmbedding g) : - ClosedEmbedding (g ∘ f) ↔ ClosedEmbedding f := by - simp_rw [closedEmbedding_iff, hg.toEmbedding.of_comp_iff, Set.range_comp, +theorem of_comp_iff (hg : IsClosedEmbedding g) : + IsClosedEmbedding (g ∘ f) ↔ IsClosedEmbedding f := by + simp_rw [isClosedEmbedding_iff, hg.toEmbedding.of_comp_iff, Set.range_comp, ← hg.closed_iff_image_closed] -theorem closure_image_eq (hf : ClosedEmbedding f) (s : Set X) : +theorem closure_image_eq (hf : IsClosedEmbedding f) (s : Set X) : closure (f '' s) = f '' closure s := hf.isClosedMap.closure_image_eq_of_continuous hf.continuous s -end ClosedEmbedding +end IsClosedEmbedding -end ClosedEmbedding +end IsClosedEmbedding diff --git a/Mathlib/Topology/Maps/Proper/Basic.lean b/Mathlib/Topology/Maps/Proper/Basic.lean index aadee505b3cac..bd8b9f6139ec2 100644 --- a/Mathlib/Topology/Maps/Proper/Basic.lean +++ b/Mathlib/Topology/Maps/Proper/Basic.lean @@ -273,17 +273,25 @@ protected lemma IsHomeomorph.isProperMap (hf : IsHomeomorph f) : IsProperMap f : @[simp] lemma isProperMap_id : IsProperMap (id : X → X) := IsHomeomorph.id.isProperMap /-- A closed embedding is proper. -/ -lemma isProperMap_of_closedEmbedding (hf : ClosedEmbedding f) : IsProperMap f := +lemma IsClosedEmbedding.isProperMap (hf : IsClosedEmbedding f) : IsProperMap f := isProperMap_of_isClosedMap_of_inj hf.continuous hf.inj hf.isClosedMap +@[deprecated (since := "2024-10-20")] +alias isProperMap_of_closedEmbedding := IsClosedEmbedding.isProperMap + /-- The coercion from a closed subset is proper. -/ -lemma isProperMap_subtype_val_of_closed {U : Set X} (hU : IsClosed U) : IsProperMap ((↑) : U → X) := - isProperMap_of_closedEmbedding hU.closedEmbedding_subtype_val +lemma IsClosed.isProperMap_subtypeVal {C : Set X} (hC : IsClosed C) : IsProperMap ((↑) : C → X) := + hC.isClosedEmbedding_subtypeVal.isProperMap + +@[deprecated (since := "2024-10-20")] +alias isProperMap_subtype_val_of_closed := IsClosed.isProperMap_subtypeVal /-- The restriction of a proper map to a closed subset is proper. -/ -lemma isProperMap_restr_of_proper_of_closed {U : Set X} (hf : IsProperMap f) (hU : IsClosed U) : - IsProperMap (fun x : U ↦ f x) := - IsProperMap.comp (isProperMap_subtype_val_of_closed hU) hf +lemma IsProperMap.restrict {C : Set X} (hf : IsProperMap f) (hC : IsClosed C) : + IsProperMap fun x : C ↦ f x := hC.isProperMap_subtypeVal.comp hf + +@[deprecated (since := "2024-10-20")] +alias isProperMap_restr_of_proper_of_closed := IsProperMap.restrict /-- The range of a proper map is closed. -/ lemma IsProperMap.isClosed_range (hf : IsProperMap f) : IsClosed (range f) := diff --git a/Mathlib/Topology/MetricSpace/Antilipschitz.lean b/Mathlib/Topology/MetricSpace/Antilipschitz.lean index 67c88ebcb3216..3ee1290ac0f7c 100644 --- a/Mathlib/Topology/MetricSpace/Antilipschitz.lean +++ b/Mathlib/Topology/MetricSpace/Antilipschitz.lean @@ -165,11 +165,14 @@ theorem isClosed_range {α β : Type*} [PseudoEMetricSpace α] [EMetricSpace β] IsClosed (range f) := (hf.isComplete_range hfc).isClosed -theorem closedEmbedding {α : Type*} {β : Type*} [EMetricSpace α] [EMetricSpace β] {K : ℝ≥0} +theorem isClosedEmbedding {α : Type*} {β : Type*} [EMetricSpace α] [EMetricSpace β] {K : ℝ≥0} {f : α → β} [CompleteSpace α] (hf : AntilipschitzWith K f) (hfc : UniformContinuous f) : - ClosedEmbedding f := + IsClosedEmbedding f := { (hf.isUniformEmbedding hfc).embedding with isClosed_range := hf.isClosed_range hfc } +@[deprecated (since := "2024-10-20")] +alias closedEmbedding := isClosedEmbedding + theorem subtype_coe (s : Set α) : AntilipschitzWith 1 ((↑) : s → α) := AntilipschitzWith.id.restrict s diff --git a/Mathlib/Topology/MetricSpace/Basic.lean b/Mathlib/Topology/MetricSpace/Basic.lean index 185ec4912deba..884f0a045a2a0 100644 --- a/Mathlib/Topology/MetricSpace/Basic.lean +++ b/Mathlib/Topology/MetricSpace/Basic.lean @@ -56,10 +56,13 @@ theorem isClosed_of_pairwise_le_dist {s : Set γ} {ε : ℝ} (hε : 0 < ε) (hs : s.Pairwise fun x y => ε ≤ dist x y) : IsClosed s := isClosed_of_spaced_out (dist_mem_uniformity hε) <| by simpa using hs -theorem closedEmbedding_of_pairwise_le_dist {α : Type*} [TopologicalSpace α] [DiscreteTopology α] +theorem isClosedEmbedding_of_pairwise_le_dist {α : Type*} [TopologicalSpace α] [DiscreteTopology α] {ε : ℝ} (hε : 0 < ε) {f : α → γ} (hf : Pairwise fun x y => ε ≤ dist (f x) (f y)) : - ClosedEmbedding f := - closedEmbedding_of_spaced_out (dist_mem_uniformity hε) <| by simpa using hf + IsClosedEmbedding f := + isClosedEmbedding_of_spaced_out (dist_mem_uniformity hε) <| by simpa using hf + +@[deprecated (since := "2024-10-20")] +alias closedEmbedding_of_pairwise_le_dist := isClosedEmbedding_of_pairwise_le_dist /-- If `f : β → α` sends any two distinct points to points at distance at least `ε > 0`, then `f` is a uniform embedding with respect to the discrete uniformity on `β`. -/ diff --git a/Mathlib/Topology/MetricSpace/Dilation.lean b/Mathlib/Topology/MetricSpace/Dilation.lean index 0ece9c35abd15..a0d5113e469b3 100644 --- a/Mathlib/Topology/MetricSpace/Dilation.lean +++ b/Mathlib/Topology/MetricSpace/Dilation.lean @@ -432,9 +432,11 @@ protected theorem embedding [PseudoEMetricSpace β] [DilationClass F α β] (f : (Dilation.isUniformEmbedding f).embedding /-- A dilation from a complete emetric space is a closed embedding -/ -protected theorem closedEmbedding [CompleteSpace α] [EMetricSpace β] [DilationClass F α β] (f : F) : - ClosedEmbedding f := - (antilipschitz f).closedEmbedding (lipschitz f).uniformContinuous +lemma isClosedEmbedding [CompleteSpace α] [EMetricSpace β] [DilationClass F α β] (f : F) : + IsClosedEmbedding f := + (antilipschitz f).isClosedEmbedding (lipschitz f).uniformContinuous + +@[deprecated (since := "2024-10-20")] alias closedEmbedding := isClosedEmbedding end EmetricDilation diff --git a/Mathlib/Topology/MetricSpace/Isometry.lean b/Mathlib/Topology/MetricSpace/Isometry.lean index 34b17ebabb3e4..21177706ee722 100644 --- a/Mathlib/Topology/MetricSpace/Isometry.lean +++ b/Mathlib/Topology/MetricSpace/Isometry.lean @@ -177,9 +177,12 @@ protected theorem embedding (hf : Isometry f) : Embedding f := hf.isUniformEmbedding.embedding /-- An isometry from a complete emetric space is a closed embedding -/ -theorem closedEmbedding [CompleteSpace α] [EMetricSpace γ] {f : α → γ} (hf : Isometry f) : - ClosedEmbedding f := - hf.antilipschitz.closedEmbedding hf.lipschitz.uniformContinuous +theorem isClosedEmbedding [CompleteSpace α] [EMetricSpace γ] {f : α → γ} (hf : Isometry f) : + IsClosedEmbedding f := + hf.antilipschitz.isClosedEmbedding hf.lipschitz.uniformContinuous + +@[deprecated (since := "2024-10-20")] +alias closedEmbedding := isClosedEmbedding end EmetricIsometry diff --git a/Mathlib/Topology/MetricSpace/Polish.lean b/Mathlib/Topology/MetricSpace/Polish.lean index 6db429aedccbf..c1e6a80a0300c 100644 --- a/Mathlib/Topology/MetricSpace/Polish.lean +++ b/Mathlib/Topology/MetricSpace/Polish.lean @@ -140,8 +140,8 @@ theorem exists_nat_nat_continuous_surjective (α : Type*) [TopologicalSpace α] exists_nat_nat_continuous_surjective_of_completeSpace α /-- Given a closed embedding into a Polish space, the source space is also Polish. -/ -theorem _root_.ClosedEmbedding.polishSpace [TopologicalSpace α] [TopologicalSpace β] [PolishSpace β] - {f : α → β} (hf : ClosedEmbedding f) : PolishSpace α := by +lemma _root_.IsClosedEmbedding.polishSpace [TopologicalSpace α] [TopologicalSpace β] [PolishSpace β] + {f : α → β} (hf : IsClosedEmbedding f) : PolishSpace α := by letI := upgradePolishSpace β letI : MetricSpace α := hf.toEmbedding.comapMetricSpace f haveI : SecondCountableTopology α := hf.toEmbedding.secondCountableTopology @@ -150,25 +150,28 @@ theorem _root_.ClosedEmbedding.polishSpace [TopologicalSpace α] [TopologicalSpa exact hf.isClosed_range.isComplete infer_instance +@[deprecated (since := "2024-10-20")] +alias _root_.ClosedEmbedding.polishSpace := _root_.IsClosedEmbedding.polishSpace + /-- Any countable discrete space is Polish. -/ instance (priority := 50) polish_of_countable [TopologicalSpace α] [h : Countable α] [DiscreteTopology α] : PolishSpace α := by obtain ⟨f, hf⟩ := h.exists_injective_nat - have : ClosedEmbedding f := by - apply closedEmbedding_of_continuous_injective_closed continuous_of_discreteTopology hf - exact fun t _ => isClosed_discrete _ + have : IsClosedEmbedding f := + .of_continuous_injective_isClosedMap continuous_of_discreteTopology hf + fun t _ ↦ isClosed_discrete _ exact this.polishSpace /-- Pulling back a Polish topology under an equiv gives again a Polish topology. -/ theorem _root_.Equiv.polishSpace_induced [t : TopologicalSpace β] [PolishSpace β] (f : α ≃ β) : @PolishSpace α (t.induced f) := letI : TopologicalSpace α := t.induced f - (f.toHomeomorphOfInducing ⟨rfl⟩).closedEmbedding.polishSpace + (f.toHomeomorphOfInducing ⟨rfl⟩).isClosedEmbedding.polishSpace /-- A closed subset of a Polish space is also Polish. -/ theorem _root_.IsClosed.polishSpace [TopologicalSpace α] [PolishSpace α] {s : Set α} (hs : IsClosed s) : PolishSpace s := - (IsClosed.closedEmbedding_subtype_val hs).polishSpace + hs.isClosedEmbedding_subtypeVal.polishSpace instance instPolishSpaceUniv [TopologicalSpace α] [PolishSpace α] : PolishSpace (univ : Set α) := @@ -211,7 +214,7 @@ theorem exists_polishSpace_forall_le {ι : Type*} [Countable ι] [t : Topologica .iInf ⟨none, Option.forall.2 ⟨le_rfl, hm⟩⟩ <| Option.forall.2 ⟨p, h'm⟩⟩ instance : PolishSpace ENNReal := - ClosedEmbedding.polishSpace ⟨ENNReal.orderIsoUnitIntervalBirational.toHomeomorph.embedding, + IsClosedEmbedding.polishSpace ⟨ENNReal.orderIsoUnitIntervalBirational.toHomeomorph.embedding, ENNReal.orderIsoUnitIntervalBirational.range_eq ▸ isClosed_univ⟩ end PolishSpace diff --git a/Mathlib/Topology/SeparatedMap.lean b/Mathlib/Topology/SeparatedMap.lean index ccb0ea0f688df..dae896a7a7e4b 100644 --- a/Mathlib/Topology/SeparatedMap.lean +++ b/Mathlib/Topology/SeparatedMap.lean @@ -86,15 +86,18 @@ theorem isSeparatedMap_iff_isClosed_diagonal {f : X → Y} : · obtain ⟨s₁, h₁, s₂, h₂, s_sub⟩ := mem_prod_iff.mp ht exact ⟨s₁, h₁, s₂, h₂, disjoint_left.2 fun x h₁ h₂ ↦ @t_sub ⟨(x, x), rfl⟩ (s_sub ⟨h₁, h₂⟩) rfl⟩ -theorem isSeparatedMap_iff_closedEmbedding {f : X → Y} : - IsSeparatedMap f ↔ ClosedEmbedding (toPullbackDiag f) := by +theorem isSeparatedMap_iff_isClosedEmbedding {f : X → Y} : + IsSeparatedMap f ↔ IsClosedEmbedding (toPullbackDiag f) := by rw [isSeparatedMap_iff_isClosed_diagonal, ← range_toPullbackDiag] exact ⟨fun h ↦ ⟨embedding_toPullbackDiag f, h⟩, fun h ↦ h.isClosed_range⟩ +@[deprecated (since := "2024-10-20")] +alias isSeparatedMap_iff_closedEmbedding := isSeparatedMap_iff_isClosedEmbedding + theorem isSeparatedMap_iff_isClosedMap {f : X → Y} : IsSeparatedMap f ↔ IsClosedMap (toPullbackDiag f) := - isSeparatedMap_iff_closedEmbedding.trans - ⟨ClosedEmbedding.isClosedMap, closedEmbedding_of_continuous_injective_closed + isSeparatedMap_iff_isClosedEmbedding.trans + ⟨IsClosedEmbedding.isClosedMap, .of_continuous_injective_isClosedMap (embedding_toPullbackDiag f).continuous (injective_toPullbackDiag f)⟩ open Function.Pullback in diff --git a/Mathlib/Topology/Separation.lean b/Mathlib/Topology/Separation.lean index d2e132acfed38..386aaf8de79c8 100644 --- a/Mathlib/Topology/Separation.lean +++ b/Mathlib/Topology/Separation.lean @@ -999,17 +999,18 @@ theorem disjoint_nhdsWithin_of_mem_discrete {s : Set X} [DiscreteTopology s] {x ⟨{x}ᶜ ∩ V, inter_mem_nhdsWithin _ h, disjoint_iff_inter_eq_empty.mpr (by rw [inter_assoc, h', compl_inter_self])⟩ -theorem closedEmbedding_update {ι : Type*} {β : ι → Type*} +theorem isClosedEmbedding_update {ι : Type*} {β : ι → Type*} [DecidableEq ι] [(i : ι) → TopologicalSpace (β i)] (x : (i : ι) → β i) (i : ι) [(i : ι) → T1Space (β i)] : - ClosedEmbedding (update x i) := by - apply closedEmbedding_of_continuous_injective_closed - · exact continuous_const.update i continuous_id - · exact update_injective x i - · intro s hs - rw [update_image] - apply isClosed_set_pi - simp [forall_update_iff, hs, isClosed_singleton] + IsClosedEmbedding (update x i) := by + refine .of_continuous_injective_isClosedMap (continuous_const.update i continuous_id) + (update_injective x i) fun s hs ↦ ?_ + rw [update_image] + apply isClosed_set_pi + simp [forall_update_iff, hs, isClosed_singleton] + +@[deprecated (since := "2024-10-20")] +alias closedEmbedding_update := isClosedEmbedding_update /-! ### R₁ (preregular) spaces -/ @@ -1754,10 +1755,13 @@ theorem Function.LeftInverse.isClosed_range [T2Space X] {f : X → Y} {g : Y → @[deprecated (since := "2024-03-17")] alias Function.LeftInverse.closed_range := Function.LeftInverse.isClosed_range -theorem Function.LeftInverse.closedEmbedding [T2Space X] {f : X → Y} {g : Y → X} - (h : Function.LeftInverse f g) (hf : Continuous f) (hg : Continuous g) : ClosedEmbedding g := +theorem Function.LeftInverse.isClosedEmbedding [T2Space X] {f : X → Y} {g : Y → X} + (h : Function.LeftInverse f g) (hf : Continuous f) (hg : Continuous g) : IsClosedEmbedding g := ⟨h.embedding hf hg, h.isClosed_range hf hg⟩ +@[deprecated (since := "2024-10-20")] +alias Function.LeftInverse.closedEmbedding := Function.LeftInverse.isClosedEmbedding + theorem SeparatedNhds.of_isCompact_isCompact [T2Space X] {s t : Set X} (hs : IsCompact s) (ht : IsCompact t) (hst : Disjoint s t) : SeparatedNhds s t := by simp only [SeparatedNhds, prod_subset_compl_diagonal_iff_disjoint.symm] at hst ⊢ @@ -1852,9 +1856,12 @@ protected theorem Continuous.isClosedMap [CompactSpace X] [T2Space Y] {f : X → (h : Continuous f) : IsClosedMap f := fun _s hs => (hs.isCompact.image h).isClosed /-- A continuous injective map from a compact space to a Hausdorff space is a closed embedding. -/ -theorem Continuous.closedEmbedding [CompactSpace X] [T2Space Y] {f : X → Y} (h : Continuous f) - (hf : Function.Injective f) : ClosedEmbedding f := - closedEmbedding_of_continuous_injective_closed h hf h.isClosedMap +theorem Continuous.isClosedEmbedding [CompactSpace X] [T2Space Y] {f : X → Y} (h : Continuous f) + (hf : Function.Injective f) : IsClosedEmbedding f := + .of_continuous_injective_isClosedMap h hf h.isClosedMap + +@[deprecated (since := "2024-10-20")] +alias Continuous.closedEmbedding := Continuous.isClosedEmbedding /-- A continuous surjective map from a compact space to a Hausdorff space is a quotient map. -/ theorem QuotientMap.of_surjective_continuous [CompactSpace X] [T2Space Y] {f : X → Y} @@ -2257,14 +2264,17 @@ theorem normal_exists_closure_subset [NormalSpace X] {s t : Set X} (hs : IsClose exact fun x hxs hxt => hs't'.le_bot ⟨hxs, hxt⟩ /-- If the codomain of a closed embedding is a normal space, then so is the domain. -/ -protected theorem ClosedEmbedding.normalSpace [TopologicalSpace Y] [NormalSpace Y] {f : X → Y} - (hf : ClosedEmbedding f) : NormalSpace X where +protected theorem IsClosedEmbedding.normalSpace [TopologicalSpace Y] [NormalSpace Y] {f : X → Y} + (hf : IsClosedEmbedding f) : NormalSpace X where normal s t hs ht hst := by have H : SeparatedNhds (f '' s) (f '' t) := NormalSpace.normal (f '' s) (f '' t) (hf.isClosedMap s hs) (hf.isClosedMap t ht) (disjoint_image_of_injective hf.inj hst) exact (H.preimage hf.continuous).mono (subset_preimage_image _ _) (subset_preimage_image _ _) +@[deprecated (since := "2024-10-20")] +alias ClosedEmbedding.normalSpace := IsClosedEmbedding.normalSpace + instance (priority := 100) NormalSpace.of_compactSpace_r1Space [CompactSpace X] [R1Space X] : NormalSpace X where normal _s _t hs ht := .of_isCompact_isCompact_isClosed hs.isCompact ht.isCompact ht @@ -2302,13 +2312,16 @@ theorem T4Space.of_compactSpace_t2Space [CompactSpace X] [T2Space X] : T4Space X := inferInstance /-- If the codomain of a closed embedding is a T₄ space, then so is the domain. -/ -protected theorem ClosedEmbedding.t4Space [TopologicalSpace Y] [T4Space Y] {f : X → Y} - (hf : ClosedEmbedding f) : T4Space X where +protected theorem IsClosedEmbedding.t4Space [TopologicalSpace Y] [T4Space Y] {f : X → Y} + (hf : IsClosedEmbedding f) : T4Space X where toT1Space := hf.toEmbedding.t1Space toNormalSpace := hf.normalSpace +@[deprecated (since := "2024-10-20")] +alias ClosedEmbedding.t4Space := IsClosedEmbedding.t4Space + instance ULift.instT4Space [T4Space X] : T4Space (ULift X) := - ULift.closedEmbedding_down.t4Space + ULift.isClosedEmbedding_down.t4Space namespace SeparationQuotient @@ -2596,7 +2609,7 @@ theorem loc_compact_Haus_tot_disc_of_zero_dim [TotallyDisconnectedSpace H] : haveI : CompactSpace s := isCompact_iff_compactSpace.1 comp obtain ⟨V : Set s, VisClopen, Vx, V_sub⟩ := compact_exists_isClopen_in_isOpen u_open_in_s xs have VisClopen' : IsClopen (((↑) : s → H) '' V) := by - refine ⟨comp.isClosed.closedEmbedding_subtype_val.closed_iff_image_closed.1 VisClopen.1, ?_⟩ + refine ⟨comp.isClosed.isClosedEmbedding_subtypeVal.closed_iff_image_closed.1 VisClopen.1, ?_⟩ let v : Set u := ((↑) : u → s) ⁻¹' V have : ((↑) : u → H) = ((↑) : s → H) ∘ ((↑) : u → s) := rfl have f0 : Embedding ((↑) : u → H) := embedding_subtype_val.comp embedding_subtype_val diff --git a/Mathlib/Topology/Sober.lean b/Mathlib/Topology/Sober.lean index 06f9286a3e5d3..6cf96933116ef 100644 --- a/Mathlib/Topology/Sober.lean +++ b/Mathlib/Topology/Sober.lean @@ -172,7 +172,7 @@ noncomputable def irreducibleSetEquivPoints [QuasiSober α] [T0Space α] : simp [hs'.closure_eq, ht'.closure_eq] rfl -theorem ClosedEmbedding.quasiSober {f : α → β} (hf : ClosedEmbedding f) [QuasiSober β] : +theorem IsClosedEmbedding.quasiSober {f : α → β} (hf : IsClosedEmbedding f) [QuasiSober β] : QuasiSober α where sober hS hS' := by have hS'' := hS.image f hf.continuous.continuousOn @@ -182,6 +182,9 @@ theorem ClosedEmbedding.quasiSober {f : α → β} (hf : ClosedEmbedding f) [Qua apply image_injective.mpr hf.inj rw [← hx.def, ← hf.closure_image_eq, image_singleton] +@[deprecated (since := "2024-10-20")] +alias ClosedEmbedding.quasiSober := IsClosedEmbedding.quasiSober + theorem IsOpenEmbedding.quasiSober {f : α → β} (hf : IsOpenEmbedding f) [QuasiSober β] : QuasiSober α where sober hS hS' := by diff --git a/Mathlib/Topology/Support.lean b/Mathlib/Topology/Support.lean index 11a71da0311a5..f0d049776a387 100644 --- a/Mathlib/Topology/Support.lean +++ b/Mathlib/Topology/Support.lean @@ -201,13 +201,16 @@ theorem _root_.hasCompactMulSupport_comp_left (hg : ∀ {x}, g x = 1 ↔ x = 1) simp_rw [hasCompactMulSupport_def, mulSupport_comp_eq g (@hg) f] @[to_additive] -theorem comp_closedEmbedding (hf : HasCompactMulSupport f) {g : α' → α} - (hg : ClosedEmbedding g) : HasCompactMulSupport (f ∘ g) := by +theorem comp_isClosedEmbedding (hf : HasCompactMulSupport f) {g : α' → α} + (hg : IsClosedEmbedding g) : HasCompactMulSupport (f ∘ g) := by rw [hasCompactMulSupport_def, Function.mulSupport_comp_eq_preimage] refine IsCompact.of_isClosed_subset (hg.isCompact_preimage hf) isClosed_closure ?_ rw [hg.toEmbedding.closure_eq_preimage_closure_image] exact preimage_mono (closure_mono <| image_preimage_subset _ _) +@[deprecated (since := "2024-10-20")] +alias comp_closedEmbedding := comp_isClosedEmbedding + @[to_additive] theorem comp₂_left (hf : HasCompactMulSupport f) (hf₂ : HasCompactMulSupport f₂) (hm : m 1 1 = 1) : diff --git a/Mathlib/Topology/TietzeExtension.lean b/Mathlib/Topology/TietzeExtension.lean index 9a2271b7f1202..70735e4ee49a4 100644 --- a/Mathlib/Topology/TietzeExtension.lean +++ b/Mathlib/Topology/TietzeExtension.lean @@ -20,7 +20,7 @@ bounded function, then there exists a bounded extension of the same norm. The proof mostly follows . We patch a small gap in the proof for unbounded functions, see -`exists_extension_forall_exists_le_ge_of_closedEmbedding`. +`exists_extension_forall_exists_le_ge_of_isClosedEmbedding`. In addition we provide a class `TietzeExtension` encoding the idea that a topological space satisfies the Tietze extension theorem. This allows us to get a version of the Tietze extension @@ -68,7 +68,7 @@ theorem ContinuousMap.exists_restrict_eq (hs : IsClosed s) (f : C(s, Y)) : nonempty topological space `X₁` into a normal topological space `X`. Let `f` be a continuous function on `X₁` with values in a `TietzeExtension` space `Y`. Then there exists a continuous function `g : C(X, Y)` such that `g ∘ e = f`. -/ -theorem ContinuousMap.exists_extension (he : ClosedEmbedding e) (f : C(X₁, Y)) : +theorem ContinuousMap.exists_extension (he : IsClosedEmbedding e) (f : C(X₁, Y)) : ∃ (g : C(X, Y)), g.comp ⟨e, he.continuous⟩ = f := by let e' : X₁ ≃ₜ Set.range e := Homeomorph.ofEmbedding _ he.toEmbedding obtain ⟨g, hg⟩ := (f.comp e'.symm).exists_restrict_eq he.isClosed_range @@ -81,7 +81,7 @@ continuous function `g : C(X, Y)` such that `g ∘ e = f`. This version is provided for convenience and backwards compatibility. Here the composition is phrased in terms of bare functions. -/ -theorem ContinuousMap.exists_extension' (he : ClosedEmbedding e) (f : C(X₁, Y)) : +theorem ContinuousMap.exists_extension' (he : IsClosedEmbedding e) (f : C(X₁, Y)) : ∃ (g : C(X, Y)), g ∘ e = f := f.exists_extension he |>.imp fun g hg ↦ by ext x; congrm($(hg) x) @@ -104,7 +104,7 @@ the radius is strictly positive, so finding this as an instance will fail. Instead, it is intended to be used as a constructor for theorems about sets which *do* satisfy `[TietzeExtension t]` under some hypotheses. -/ -theorem ContinuousMap.exists_extension_forall_mem (he : ClosedEmbedding e) +theorem ContinuousMap.exists_extension_forall_mem (he : IsClosedEmbedding e) {Y : Type v} [TopologicalSpace Y] (f : C(X₁, Y)) {t : Set Y} (hf : ∀ x, f x ∈ t) [ht : TietzeExtension.{u, v} t] : ∃ (g : C(X, Y)), (∀ x, g x ∈ t) ∧ g.comp ⟨e, he.continuous⟩ = f := by @@ -166,7 +166,7 @@ namespace BoundedContinuousFunction of a topological space into a normal topological space and `f : X →ᵇ ℝ` is a bounded continuous function, then there exists a bounded continuous function `g : Y →ᵇ ℝ` of the norm `‖g‖ ≤ ‖f‖ / 3` such that the distance between `g ∘ e` and `f` is at most `(2 / 3) * ‖f‖`. -/ -theorem tietze_extension_step (f : X →ᵇ ℝ) (e : C(X, Y)) (he : ClosedEmbedding e) : +theorem tietze_extension_step (f : X →ᵇ ℝ) (e : C(X, Y)) (he : IsClosedEmbedding e) : ∃ g : Y →ᵇ ℝ, ‖g‖ ≤ ‖f‖ / 3 ∧ dist (g.compContinuous e) f ≤ 2 / 3 * ‖f‖ := by have h3 : (0 : ℝ) < 3 := by norm_num1 have h23 : 0 < (2 / 3 : ℝ) := by norm_num1 @@ -216,8 +216,8 @@ theorem tietze_extension_step (f : X →ᵇ ℝ) (e : C(X, Y)) (he : ClosedEmbed embedding and bundled composition. If `e : C(X, Y)` is a closed embedding of a topological space into a normal topological space and `f : X →ᵇ ℝ` is a bounded continuous function, then there exists a bounded continuous function `g : Y →ᵇ ℝ` of the same norm such that `g ∘ e = f`. -/ -theorem exists_extension_norm_eq_of_closedEmbedding' (f : X →ᵇ ℝ) (e : C(X, Y)) - (he : ClosedEmbedding e) : ∃ g : Y →ᵇ ℝ, ‖g‖ = ‖f‖ ∧ g.compContinuous e = f := by +theorem exists_extension_norm_eq_of_isClosedEmbedding' (f : X →ᵇ ℝ) (e : C(X, Y)) + (he : IsClosedEmbedding e) : ∃ g : Y →ᵇ ℝ, ‖g‖ = ‖f‖ ∧ g.compContinuous e = f := by /- For the proof, we iterate `tietze_extension_step`. Each time we apply it to the difference between the previous approximation and `f`. -/ choose F hF_norm hF_dist using fun f : X →ᵇ ℝ => tietze_extension_step f e he @@ -261,36 +261,42 @@ theorem exists_extension_norm_eq_of_closedEmbedding' (f : X →ᵇ ℝ) (e : C(X · rw [← hge] exact norm_compContinuous_le _ _ +@[deprecated (since := "2024-10-20")] +alias exists_extension_norm_eq_of_closedEmbedding' := exists_extension_norm_eq_of_isClosedEmbedding' + /-- **Tietze extension theorem** for real-valued bounded continuous maps, a version with a closed embedding and unbundled composition. If `e : C(X, Y)` is a closed embedding of a topological space into a normal topological space and `f : X →ᵇ ℝ` is a bounded continuous function, then there exists a bounded continuous function `g : Y →ᵇ ℝ` of the same norm such that `g ∘ e = f`. -/ -theorem exists_extension_norm_eq_of_closedEmbedding (f : X →ᵇ ℝ) {e : X → Y} - (he : ClosedEmbedding e) : ∃ g : Y →ᵇ ℝ, ‖g‖ = ‖f‖ ∧ g ∘ e = f := by - rcases exists_extension_norm_eq_of_closedEmbedding' f ⟨e, he.continuous⟩ he with ⟨g, hg, rfl⟩ +theorem exists_extension_norm_eq_of_isClosedEmbedding (f : X →ᵇ ℝ) {e : X → Y} + (he : IsClosedEmbedding e) : ∃ g : Y →ᵇ ℝ, ‖g‖ = ‖f‖ ∧ g ∘ e = f := by + rcases exists_extension_norm_eq_of_isClosedEmbedding' f ⟨e, he.continuous⟩ he with ⟨g, hg, rfl⟩ exact ⟨g, hg, rfl⟩ +@[deprecated (since := "2024-10-20")] +alias exists_extension_norm_eq_of_closedEmbedding := exists_extension_norm_eq_of_isClosedEmbedding + /-- **Tietze extension theorem** for real-valued bounded continuous maps, a version for a closed set. If `f` is a bounded continuous real-valued function defined on a closed set in a normal topological space, then it can be extended to a bounded continuous function of the same norm defined on the whole space. -/ theorem exists_norm_eq_restrict_eq_of_closed {s : Set Y} (f : s →ᵇ ℝ) (hs : IsClosed s) : ∃ g : Y →ᵇ ℝ, ‖g‖ = ‖f‖ ∧ g.restrict s = f := - exists_extension_norm_eq_of_closedEmbedding' f ((ContinuousMap.id _).restrict s) - (closedEmbedding_subtype_val hs) + exists_extension_norm_eq_of_isClosedEmbedding' f ((ContinuousMap.id _).restrict s) + hs.isClosedEmbedding_subtypeVal /-- **Tietze extension theorem** for real-valued bounded continuous maps, a version for a closed embedding and a bounded continuous function that takes values in a non-trivial closed interval. -See also `exists_extension_forall_mem_of_closedEmbedding` for a more general statement that works +See also `exists_extension_forall_mem_of_isClosedEmbedding` for a more general statement that works for any interval (finite or infinite, open or closed). If `e : X → Y` is a closed embedding and `f : X →ᵇ ℝ` is a bounded continuous function such that `f x ∈ [a, b]` for all `x`, where `a ≤ b`, then there exists a bounded continuous function `g : Y →ᵇ ℝ` such that `g y ∈ [a, b]` for all `y` and `g ∘ e = f`. -/ -theorem exists_extension_forall_mem_Icc_of_closedEmbedding (f : X →ᵇ ℝ) {a b : ℝ} {e : X → Y} - (hf : ∀ x, f x ∈ Icc a b) (hle : a ≤ b) (he : ClosedEmbedding e) : +theorem exists_extension_forall_mem_Icc_of_isClosedEmbedding (f : X →ᵇ ℝ) {a b : ℝ} {e : X → Y} + (hf : ∀ x, f x ∈ Icc a b) (hle : a ≤ b) (he : IsClosedEmbedding e) : ∃ g : Y →ᵇ ℝ, (∀ y, g y ∈ Icc a b) ∧ g ∘ e = f := by - rcases exists_extension_norm_eq_of_closedEmbedding (f - const X ((a + b) / 2)) he with + rcases exists_extension_norm_eq_of_isClosedEmbedding (f - const X ((a + b) / 2)) he with ⟨g, hgf, hge⟩ refine ⟨const Y ((a + b) / 2) + g, fun y => ?_, ?_⟩ · suffices ‖f - const X ((a + b) / 2)‖ ≤ (b - a) / 2 by @@ -302,13 +308,17 @@ theorem exists_extension_forall_mem_Icc_of_closedEmbedding (f : X →ᵇ ℝ) {a have : g (e x) = f x - (a + b) / 2 := congr_fun hge x simp [this] +@[deprecated (since := "2024-10-20")] +alias exists_extension_forall_mem_Icc_of_closedEmbedding := + exists_extension_forall_mem_Icc_of_isClosedEmbedding + /-- **Tietze extension theorem** for real-valued bounded continuous maps, a version for a closed embedding. Let `e` be a closed embedding of a nonempty topological space `X` into a normal topological space `Y`. Let `f` be a bounded continuous real-valued function on `X`. Then there exists a bounded continuous function `g : Y →ᵇ ℝ` such that `g ∘ e = f` and each value `g y` belongs to a closed interval `[f x₁, f x₂]` for some `x₁` and `x₂`. -/ -theorem exists_extension_forall_exists_le_ge_of_closedEmbedding [Nonempty X] (f : X →ᵇ ℝ) - {e : X → Y} (he : ClosedEmbedding e) : +theorem exists_extension_forall_exists_le_ge_of_isClosedEmbedding [Nonempty X] (f : X →ᵇ ℝ) + {e : X → Y} (he : IsClosedEmbedding e) : ∃ g : Y →ᵇ ℝ, (∀ y, ∃ x₁ x₂, g y ∈ Icc (f x₁) (f x₂)) ∧ g ∘ e = f := by inhabit X -- Put `a = ⨅ x, f x` and `b = ⨆ x, f x` @@ -329,12 +339,12 @@ theorem exists_extension_forall_exists_le_ge_of_closedEmbedding [Nonempty X] (f have hsub : c - a = b - c := by field_simp [c] ring - /- Due to `exists_extension_forall_mem_Icc_of_closedEmbedding`, there exists an extension `g` + /- Due to `exists_extension_forall_mem_Icc_of_isClosedEmbedding`, there exists an extension `g` such that `g y ∈ [a, b]` for all `y`. However, if `a` and/or `b` do not belong to the range of `f`, then we need to ensure that these points do not belong to the range of `g`. This is done in two almost identical steps. First we deal with the case `∀ x, f x ≠ a`. -/ obtain ⟨g, hg_mem, hgf⟩ : ∃ g : Y →ᵇ ℝ, (∀ y, ∃ x, g y ∈ Icc (f x) b) ∧ g ∘ e = f := by - rcases exists_extension_forall_mem_Icc_of_closedEmbedding f hmem hle he with ⟨g, hg_mem, hgf⟩ + rcases exists_extension_forall_mem_Icc_of_isClosedEmbedding f hmem hle he with ⟨g, hg_mem, hgf⟩ -- If `a ∈ range f`, then we are done. rcases em (∃ x, f x = a) with (⟨x, rfl⟩ | ha') · exact ⟨g, fun y => ⟨x, hg_mem _⟩, hgf⟩ @@ -412,6 +422,10 @@ theorem exists_extension_forall_exists_le_ge_of_closedEmbedding [Nonempty X] (f · refine ⟨xl y, xu, ?_, hyxu.le⟩ simp [dg0 (Or.inr hc), hxl] +@[deprecated (since := "2024-10-20")] +alias exists_extension_forall_exists_le_ge_of_closedEmbedding := + exists_extension_forall_exists_le_ge_of_isClosedEmbedding + /-- **Tietze extension theorem** for real-valued bounded continuous maps, a version for a closed embedding. Let `e` be a closed embedding of a nonempty topological space `X` into a normal topological space `Y`. Let `f` be a bounded continuous real-valued function on `X`. Let `t` be @@ -419,17 +433,21 @@ a nonempty convex set of real numbers (we use `OrdConnected` instead of `Convex` deduce this argument by typeclass search) such that `f x ∈ t` for all `x`. Then there exists a bounded continuous real-valued function `g : Y →ᵇ ℝ` such that `g y ∈ t` for all `y` and `g ∘ e = f`. -/ -theorem exists_extension_forall_mem_of_closedEmbedding (f : X →ᵇ ℝ) {t : Set ℝ} {e : X → Y} - [hs : OrdConnected t] (hf : ∀ x, f x ∈ t) (hne : t.Nonempty) (he : ClosedEmbedding e) : +theorem exists_extension_forall_mem_of_isClosedEmbedding (f : X →ᵇ ℝ) {t : Set ℝ} {e : X → Y} + [hs : OrdConnected t] (hf : ∀ x, f x ∈ t) (hne : t.Nonempty) (he : IsClosedEmbedding e) : ∃ g : Y →ᵇ ℝ, (∀ y, g y ∈ t) ∧ g ∘ e = f := by cases isEmpty_or_nonempty X · rcases hne with ⟨c, hc⟩ exact ⟨const Y c, fun _ => hc, funext fun x => isEmptyElim x⟩ - rcases exists_extension_forall_exists_le_ge_of_closedEmbedding f he with ⟨g, hg, hgf⟩ + rcases exists_extension_forall_exists_le_ge_of_isClosedEmbedding f he with ⟨g, hg, hgf⟩ refine ⟨g, fun y => ?_, hgf⟩ rcases hg y with ⟨xl, xu, h⟩ exact hs.out (hf _) (hf _) h +@[deprecated (since := "2024-10-20")] +alias exists_extension_forall_mem_of_closedEmbedding := + exists_extension_forall_mem_of_isClosedEmbedding + /-- **Tietze extension theorem** for real-valued bounded continuous maps, a version for a closed set. Let `s` be a closed set in a normal topological space `Y`. Let `f` be a bounded continuous real-valued function on `s`. Let `t` be a nonempty convex set of real numbers (we use @@ -439,9 +457,8 @@ that `f x ∈ t` for all `x : s`. Then there exists a bounded continuous real-va theorem exists_forall_mem_restrict_eq_of_closed {s : Set Y} (f : s →ᵇ ℝ) (hs : IsClosed s) {t : Set ℝ} [OrdConnected t] (hf : ∀ x, f x ∈ t) (hne : t.Nonempty) : ∃ g : Y →ᵇ ℝ, (∀ y, g y ∈ t) ∧ g.restrict s = f := by - rcases exists_extension_forall_mem_of_closedEmbedding f hf hne - (closedEmbedding_subtype_val hs) with - ⟨g, hg, hgf⟩ + obtain ⟨g, hg, hgf⟩ := + exists_extension_forall_mem_of_isClosedEmbedding f hf hne hs.isClosedEmbedding_subtypeVal exact ⟨g, hg, DFunLike.coe_injective hgf⟩ end BoundedContinuousFunction @@ -454,8 +471,8 @@ topological space `Y`. Let `f` be a continuous real-valued function on `X`. Let convex set of real numbers (we use `OrdConnected` instead of `Convex` to automatically deduce this argument by typeclass search) such that `f x ∈ t` for all `x`. Then there exists a continuous real-valued function `g : C(Y, ℝ)` such that `g y ∈ t` for all `y` and `g ∘ e = f`. -/ -theorem exists_extension_forall_mem_of_closedEmbedding (f : C(X, ℝ)) {t : Set ℝ} {e : X → Y} - [hs : OrdConnected t] (hf : ∀ x, f x ∈ t) (hne : t.Nonempty) (he : ClosedEmbedding e) : +theorem exists_extension_forall_mem_of_isClosedEmbedding (f : C(X, ℝ)) {t : Set ℝ} {e : X → Y} + [hs : OrdConnected t] (hf : ∀ x, f x ∈ t) (hne : t.Nonempty) (he : IsClosedEmbedding e) : ∃ g : C(Y, ℝ), (∀ y, g y ∈ t) ∧ g ∘ e = f := by have h : ℝ ≃o Ioo (-1 : ℝ) 1 := orderIsoIooNegOneOne ℝ let F : X →ᵇ ℝ := @@ -474,7 +491,7 @@ theorem exists_extension_forall_mem_of_closedEmbedding (f : C(X, ℝ)) {t : Set rcases hz with ⟨z, hz, rfl⟩ exact ⟨z, hs.out hx hy hz, rfl⟩ have hFt : ∀ x, F x ∈ t' := fun x => mem_image_of_mem _ (hf x) - rcases F.exists_extension_forall_mem_of_closedEmbedding hFt (hne.image _) he with ⟨G, hG, hGF⟩ + rcases F.exists_extension_forall_mem_of_isClosedEmbedding hFt (hne.image _) he with ⟨G, hG, hGF⟩ let g : C(Y, ℝ) := ⟨h.symm ∘ codRestrict G _ fun y => ht_sub (hG y), h.symm.continuous.comp <| G.continuous.subtype_mk _⟩ @@ -487,8 +504,15 @@ theorem exists_extension_forall_mem_of_closedEmbedding (f : C(X, ℝ)) {t : Set · ext x exact hgG.2 (congr_fun hGF _) +@[deprecated (since := "2024-10-20")] +alias exists_extension_forall_mem_of_closedEmbedding := + exists_extension_forall_mem_of_isClosedEmbedding + @[deprecated (since := "2024-01-16")] -alias exists_extension_of_closedEmbedding := exists_extension' +alias exists_extension_of_isClosedEmbedding := exists_extension' + +@[deprecated (since := "2024-10-20")] +alias exists_extension_of_closedEmbedding := exists_extension_of_isClosedEmbedding /-- **Tietze extension theorem** for real-valued continuous maps, a version for a closed set. Let `s` be a closed set in a normal topological space `Y`. Let `f` be a continuous real-valued function @@ -500,7 +524,7 @@ theorem exists_restrict_eq_forall_mem_of_closed {s : Set Y} (f : C(s, ℝ)) {t : [OrdConnected t] (ht : ∀ x, f x ∈ t) (hne : t.Nonempty) (hs : IsClosed s) : ∃ g : C(Y, ℝ), (∀ y, g y ∈ t) ∧ g.restrict s = f := let ⟨g, hgt, hgf⟩ := - exists_extension_forall_mem_of_closedEmbedding f ht hne (closedEmbedding_subtype_val hs) + exists_extension_forall_mem_of_isClosedEmbedding f ht hne hs.isClosedEmbedding_subtypeVal ⟨g, hgt, coe_injective hgf⟩ @[deprecated (since := "2024-01-16")] alias exists_restrict_eq_of_closed := exists_restrict_eq diff --git a/Mathlib/Topology/UniformSpace/Ascoli.lean b/Mathlib/Topology/UniformSpace/Ascoli.lean index 351606b63c56c..4de38296cb8f6 100644 --- a/Mathlib/Topology/UniformSpace/Ascoli.lean +++ b/Mathlib/Topology/UniformSpace/Ascoli.lean @@ -456,14 +456,17 @@ and `F : ι → (X → α)`. Assume that: * For all `x`, the range of `i ↦ F i x` is contained in some fixed compact subset. Then `ι` is compact. -/ -theorem ArzelaAscoli.compactSpace_of_closedEmbedding [TopologicalSpace ι] {𝔖 : Set (Set X)} - (𝔖_compact : ∀ K ∈ 𝔖, IsCompact K) (F_clemb : ClosedEmbedding (UniformOnFun.ofFun 𝔖 ∘ F)) +theorem ArzelaAscoli.compactSpace_of_isClosedEmbedding [TopologicalSpace ι] {𝔖 : Set (Set X)} + (𝔖_compact : ∀ K ∈ 𝔖, IsCompact K) (F_clemb : IsClosedEmbedding (UniformOnFun.ofFun 𝔖 ∘ F)) (F_eqcont : ∀ K ∈ 𝔖, EquicontinuousOn F K) (F_pointwiseCompact : ∀ K ∈ 𝔖, ∀ x ∈ K, ∃ Q, IsCompact Q ∧ ∀ i, F i x ∈ Q) : CompactSpace ι := compactSpace_of_closed_inducing' 𝔖_compact F_clemb.toInducing F_clemb.isClosed_range F_eqcont F_pointwiseCompact +@[deprecated (since := "2024-10-20")] +alias ArzelaAscoli.compactSpace_of_closedEmbedding := ArzelaAscoli.compactSpace_of_isClosedEmbedding + /-- A version of the **Arzela-Ascoli theorem**. Let `X, ι` be topological spaces, `𝔖` a covering of `X` by compact subsets, `α` a T2 uniform space, @@ -474,13 +477,13 @@ Let `X, ι` be topological spaces, `𝔖` a covering of `X` by compact subsets, * For all `x ∈ ⋃₀ 𝔖`, the image of `s` under `i ↦ F i x` is contained in some fixed compact subset. Then `s` has compact closure in `ι`. -/ -theorem ArzelaAscoli.isCompact_closure_of_closedEmbedding [TopologicalSpace ι] [T2Space α] +theorem ArzelaAscoli.isCompact_closure_of_isClosedEmbedding [TopologicalSpace ι] [T2Space α] {𝔖 : Set (Set X)} (𝔖_compact : ∀ K ∈ 𝔖, IsCompact K) - (F_clemb : ClosedEmbedding (UniformOnFun.ofFun 𝔖 ∘ F)) + (F_clemb : IsClosedEmbedding (UniformOnFun.ofFun 𝔖 ∘ F)) {s : Set ι} (s_eqcont : ∀ K ∈ 𝔖, EquicontinuousOn (F ∘ ((↑) : s → ι)) K) (s_pointwiseCompact : ∀ K ∈ 𝔖, ∀ x ∈ K, ∃ Q, IsCompact Q ∧ ∀ i ∈ s, F i x ∈ Q) : IsCompact (closure s) := by - -- We apply `ArzelaAscoli.compactSpace_of_closedEmbedding` to the map + -- We apply `ArzelaAscoli.compactSpace_of_isClosedEmbedding` to the map -- `F ∘ (↑) : closure s → (X → α)`, for which all the hypotheses are easily verified. rw [isCompact_iff_compactSpace] have : ∀ K ∈ 𝔖, ∀ x ∈ K, Continuous (eval x ∘ F) := fun K hK x hx ↦ @@ -491,10 +494,14 @@ theorem ArzelaAscoli.isCompact_closure_of_closedEmbedding [TopologicalSpace ι] have cls_pointwiseCompact : ∀ K ∈ 𝔖, ∀ x ∈ K, ∃ Q, IsCompact Q ∧ ∀ i ∈ closure s, F i x ∈ Q := fun K hK x hx ↦ (s_pointwiseCompact K hK x hx).imp fun Q hQ ↦ ⟨hQ.1, closure_minimal hQ.2 <| hQ.1.isClosed.preimage (this K hK x hx)⟩ - exact ArzelaAscoli.compactSpace_of_closedEmbedding 𝔖_compact - (F_clemb.comp isClosed_closure.closedEmbedding_subtype_val) cls_eqcont + exact ArzelaAscoli.compactSpace_of_isClosedEmbedding 𝔖_compact + (F_clemb.comp isClosed_closure.isClosedEmbedding_subtypeVal) cls_eqcont fun K hK x hx ↦ (cls_pointwiseCompact K hK x hx).imp fun Q hQ ↦ ⟨hQ.1, by simpa using hQ.2⟩ +@[deprecated (since := "2024-10-20")] +alias ArzelaAscoli.isCompact_closure_of_closedEmbedding := + ArzelaAscoli.isCompact_closure_of_isClosedEmbedding + /-- A version of the **Arzela-Ascoli theorem**. If an equicontinuous family of continuous functions is compact in the pointwise topology, then it diff --git a/Mathlib/Topology/UniformSpace/CompleteSeparated.lean b/Mathlib/Topology/UniformSpace/CompleteSeparated.lean index 7444f5a467be3..cc34886c49d41 100644 --- a/Mathlib/Topology/UniformSpace/CompleteSeparated.lean +++ b/Mathlib/Topology/UniformSpace/CompleteSeparated.lean @@ -27,13 +27,19 @@ theorem IsComplete.isClosed [UniformSpace α] [T0Space α] {s : Set α} (h : IsC rcases h f this inf_le_right with ⟨y, ys, fy⟩ rwa [(tendsto_nhds_unique' ha inf_le_left fy : a = y)] -theorem IsUniformEmbedding.toClosedEmbedding [UniformSpace α] [UniformSpace β] [CompleteSpace α] +theorem IsUniformEmbedding.toIsClosedEmbedding [UniformSpace α] [UniformSpace β] [CompleteSpace α] [T0Space β] {f : α → β} (hf : IsUniformEmbedding f) : - ClosedEmbedding f := + IsClosedEmbedding f := ⟨hf.embedding, hf.isUniformInducing.isComplete_range.isClosed⟩ +@[deprecated (since := "2024-10-20")] +alias IsUniformEmbedding.toClosedEmbedding := IsUniformEmbedding.toIsClosedEmbedding + @[deprecated (since := "2024-10-01")] -alias UniformEmbedding.toClosedEmbedding := IsUniformEmbedding.toClosedEmbedding +alias UniformEmbedding.toIsClosedEmbedding := IsUniformEmbedding.toIsClosedEmbedding + +@[deprecated (since := "2024-10-20")] +alias UniformEmbedding.toClosedEmbedding := UniformEmbedding.toIsClosedEmbedding namespace IsDenseInducing diff --git a/Mathlib/Topology/UniformSpace/UniformEmbedding.lean b/Mathlib/Topology/UniformSpace/UniformEmbedding.lean index 36dee5874506e..565e1e1a1e8be 100644 --- a/Mathlib/Topology/UniformSpace/UniformEmbedding.lean +++ b/Mathlib/Topology/UniformSpace/UniformEmbedding.lean @@ -347,14 +347,17 @@ alias UniformEmbedding.isDenseEmbedding := IsUniformEmbedding.isDenseEmbedding @[deprecated (since := "2024-09-30")] alias IsUniformEmbedding.denseEmbedding := IsUniformEmbedding.isDenseEmbedding -theorem closedEmbedding_of_spaced_out {α} [TopologicalSpace α] [DiscreteTopology α] +theorem isClosedEmbedding_of_spaced_out {α} [TopologicalSpace α] [DiscreteTopology α] [T0Space β] {f : α → β} {s : Set (β × β)} (hs : s ∈ 𝓤 β) - (hf : Pairwise fun x y => (f x, f y) ∉ s) : ClosedEmbedding f := by + (hf : Pairwise fun x y => (f x, f y) ∉ s) : IsClosedEmbedding f := by rcases @DiscreteTopology.eq_bot α _ _ with rfl; let _ : UniformSpace α := ⊥ exact { (isUniformEmbedding_of_spaced_out hs hf).embedding with isClosed_range := isClosed_range_of_spaced_out hs hf } +@[deprecated (since := "2024-10-20")] +alias closedEmbedding_of_spaced_out := isClosedEmbedding_of_spaced_out + theorem closure_image_mem_nhds_of_isUniformInducing {s : Set (α × α)} {e : α → β} (b : β) (he₁ : IsUniformInducing e) (he₂ : IsDenseInducing e) (hs : s ∈ 𝓤 α) : ∃ a, closure (e '' { a' | (a, a') ∈ s }) ∈ 𝓝 b := by From 8d27fbf359865a12c7bc7ed41bfe6ede9fd372a3 Mon Sep 17 00:00:00 2001 From: Markus Himmel Date: Sun, 20 Oct 2024 16:57:21 +0000 Subject: [PATCH 205/505] feat(CategoryTheory): pull colimits out of hom functors with Yoneda on the LHS (#17440) A technical result on the way towards showing that Ind-objects are closed under filtered colimits. Co-authored-by: Markus Himmel --- Mathlib.lean | 1 + .../Limits/Preserves/Yoneda.lean | 79 +++++++++++++++++++ 2 files changed, 80 insertions(+) create mode 100644 Mathlib/CategoryTheory/Limits/Preserves/Yoneda.lean diff --git a/Mathlib.lean b/Mathlib.lean index de1b28985d7b6..ebba90f308a9c 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1682,6 +1682,7 @@ import Mathlib.CategoryTheory.Limits.Preserves.Shapes.Square import Mathlib.CategoryTheory.Limits.Preserves.Shapes.Terminal import Mathlib.CategoryTheory.Limits.Preserves.Shapes.Zero import Mathlib.CategoryTheory.Limits.Preserves.Ulift +import Mathlib.CategoryTheory.Limits.Preserves.Yoneda import Mathlib.CategoryTheory.Limits.Presheaf import Mathlib.CategoryTheory.Limits.Shapes.BinaryProducts import Mathlib.CategoryTheory.Limits.Shapes.Biproducts diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Yoneda.lean b/Mathlib/CategoryTheory/Limits/Preserves/Yoneda.lean new file mode 100644 index 0000000000000..629c22852c6a5 --- /dev/null +++ b/Mathlib/CategoryTheory/Limits/Preserves/Yoneda.lean @@ -0,0 +1,79 @@ +/- +Copyright (c) 2024 Markus Himmel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Markus Himmel +-/ +import Mathlib.CategoryTheory.Limits.Preserves.Ulift +import Mathlib.CategoryTheory.Limits.FunctorToTypes + +/-! +# Yoneda preserves certain colimits + +Given a bifunctor `F : J ⥤ Cᵒᵖ ⥤ Type v`, we prove the isomorphism +`Hom(YX, colim_j F(j, -)) ≅ colim_j Hom(YX, F(j, -))`, where `Y` is the Yoneda embedding. + +We state this in two ways. One is functorial in `X` and stated as a natural isomorphism of functors +`yoneda.op ⋙ yoneda.obj (colimit F) ≅ yoneda.op ⋙ colimit (F ⋙ yoneda)`, and from this we +deduce the more traditional preservation statement +`PreservesColimit F (coyoneda.obj (op (yoneda.obj X)))`. + +The proof combines the Yoneda lemma with the fact that colimits in functor categories are computed +pointwise. + +## See also + +There is also a relative version of this statement where `F : J ⥤ Over A` for some presheaf +`A`, see `CategoryTheory.Comma.Presheaf`. + +-/ + +universe v₁ v₂ v₃ u₁ u₂ u₃ + +namespace CategoryTheory + +open CategoryTheory.Limits Opposite + +variable {C : Type u₁} [Category.{v₁} C] + +variable {J : Type u₂} [Category.{v₂} J] [HasColimitsOfShape J (Type v₁)] + [HasColimitsOfShape J (Type (max u₁ v₁))] (F : J ⥤ Cᵒᵖ ⥤ Type v₁) + +/-- Naturally in `X`, we have `Hom(YX, colim_i Fi) ≅ colim_i Hom(YX, Fi)`. -/ +noncomputable def yonedaYonedaColimit : + yoneda.op ⋙ yoneda.obj (colimit F) ≅ yoneda.op ⋙ colimit (F ⋙ yoneda) := calc + yoneda.op ⋙ yoneda.obj (colimit F) + ≅ colimit F ⋙ uliftFunctor.{u₁} := yonedaOpCompYonedaObj (colimit F) + _ ≅ F.flip ⋙ colim ⋙ uliftFunctor.{u₁} := + isoWhiskerRight (colimitIsoFlipCompColim F) uliftFunctor.{u₁} + _ ≅ F.flip ⋙ (whiskeringRight _ _ _).obj uliftFunctor.{u₁} ⋙ colim := + isoWhiskerLeft F.flip (preservesColimitNatIso uliftFunctor.{u₁}) + _ ≅ (yoneda.op ⋙ coyoneda ⋙ (whiskeringLeft _ _ _).obj F) ⋙ colim := isoWhiskerRight + (isoWhiskerRight largeCurriedYonedaLemma.symm ((whiskeringLeft _ _ _).obj F)) colim + _ ≅ yoneda.op ⋙ colimit (F ⋙ yoneda) := + isoWhiskerLeft yoneda.op (colimitIsoFlipCompColim (F ⋙ yoneda)).symm + +theorem yonedaYonedaColimit_app_inv {X : C} : ((yonedaYonedaColimit F).app (op X)).inv = + (colimitObjIsoColimitCompEvaluation _ _).hom ≫ + (colimit.post F (coyoneda.obj (op (yoneda.obj X)))) := by + dsimp [yonedaYonedaColimit] + simp only [Category.id_comp, Iso.cancel_iso_hom_left] + apply colimit.hom_ext + intro j + rw [colimit.ι_post, ι_colimMap_assoc] + simp only [← CategoryTheory.Functor.assoc, comp_evaluation] + rw [ι_preservesColimitsIso_inv_assoc, ← Functor.map_comp_assoc] + simp only [← comp_evaluation] + rw [colimitObjIsoColimitCompEvaluation_ι_inv, whiskerLeft_app] + ext η Y f + simp [largeCurriedYonedaLemma, yonedaOpCompYonedaObj, FunctorToTypes.colimit.map_ι_apply, + map_yonedaEquiv] + +noncomputable instance {X : C} : PreservesColimit F (coyoneda.obj (op (yoneda.obj X))) := by + suffices IsIso (colimit.post F (coyoneda.obj (op (yoneda.obj X)))) from + preservesColimitOfIsIsoPost _ _ + suffices colimit.post F (coyoneda.obj (op (yoneda.obj X))) = + (colimitObjIsoColimitCompEvaluation _ _).inv ≫ ((yonedaYonedaColimit F).app (op X)).inv from + this ▸ inferInstance + rw [yonedaYonedaColimit_app_inv, Iso.inv_hom_id_assoc] + +end CategoryTheory From 03a37e0be41d95a6f229d3578b5481c50e125cf4 Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Sun, 20 Oct 2024 17:40:31 +0000 Subject: [PATCH 206/505] chore(LinearAlgebra/TensorProduct): fix duplication (#17912) This was unintentionally made in [#16776](https://github.com/leanprover-community/mathlib4/pull/16776) by myself. --- Mathlib/Algebra/Category/AlgebraCat/Monoidal.lean | 2 +- Mathlib/LinearAlgebra/TensorProduct/Basic.lean | 9 ++------- 2 files changed, 3 insertions(+), 8 deletions(-) diff --git a/Mathlib/Algebra/Category/AlgebraCat/Monoidal.lean b/Mathlib/Algebra/Category/AlgebraCat/Monoidal.lean index 2275dee475858..28e837f52f457 100644 --- a/Mathlib/Algebra/Category/AlgebraCat/Monoidal.lean +++ b/Mathlib/Algebra/Category/AlgebraCat/Monoidal.lean @@ -60,7 +60,7 @@ noncomputable instance instMonoidalCategory : MonoidalCategory (AlgebraCat.{u} R (forget₂ (AlgebraCat R) (ModuleCat R)) { μIso := fun _ _ => Iso.refl _ εIso := Iso.refl _ - associator_eq := fun _ _ _ => TensorProduct.ext₃ (fun _ _ _ => rfl) + associator_eq := fun _ _ _ => TensorProduct.ext_threefold (fun _ _ _ => rfl) leftUnitor_eq := fun _ => TensorProduct.ext' (fun _ _ => rfl) rightUnitor_eq := fun _ => TensorProduct.ext' (fun _ _ => rfl) } diff --git a/Mathlib/LinearAlgebra/TensorProduct/Basic.lean b/Mathlib/LinearAlgebra/TensorProduct/Basic.lean index 8225c569e1d95..f08d4d96958bb 100644 --- a/Mathlib/LinearAlgebra/TensorProduct/Basic.lean +++ b/Mathlib/LinearAlgebra/TensorProduct/Basic.lean @@ -526,13 +526,6 @@ theorem ext' {g h : M ⊗[R] N →ₗ[R] P} (H : ∀ x y, g (x ⊗ₜ y) = h (x TensorProduct.induction_on z (by simp_rw [LinearMap.map_zero]) H fun x y ihx ihy => by rw [g.map_add, h.map_add, ihx, ihy] -theorem ext₃ {g h : (M ⊗[R] N) ⊗[R] P →ₗ[R] Q} - (H : ∀ x y z, g (x ⊗ₜ y ⊗ₜ z) = h (x ⊗ₜ y ⊗ₜ z)) : g = h := - ext' fun x => TensorProduct.induction_on x - (fun x => by simp only [zero_tmul, map_zero]) - (fun x y => H x y) - (fun x y ihx ihy z => by rw [add_tmul, g.map_add, h.map_add, ihx, ihy]) - theorem lift.unique {g : M ⊗[R] N →ₗ[R] P} (H : ∀ x y, g (x ⊗ₜ y) = f x y) : g = lift f := ext' fun m n => by rw [H, lift.tmul] @@ -621,6 +614,8 @@ theorem ext_threefold {g h : (M ⊗[R] N) ⊗[R] P →ₗ[R] Q} ext x y z exact H x y z +@[deprecated (since := "2024-10-18")] alias ext₃ := ext_threefold + -- We'll need this one for checking the pentagon identity! theorem ext_fourfold {g h : ((M ⊗[R] N) ⊗[R] P) ⊗[R] Q →ₗ[R] S} (H : ∀ w x y z, g (w ⊗ₜ x ⊗ₜ y ⊗ₜ z) = h (w ⊗ₜ x ⊗ₜ y ⊗ₜ z)) : g = h := by From 15fb8246a353ad60012ade7c341d315e827d89d2 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sun, 20 Oct 2024 18:58:06 +0000 Subject: [PATCH 207/505] chore(Topology/Exterior): golf (#17955) Use `exterior_mono` to golf 3 lemmas --- Mathlib/Topology/Exterior.lean | 24 +++++++++--------------- 1 file changed, 9 insertions(+), 15 deletions(-) diff --git a/Mathlib/Topology/Exterior.lean b/Mathlib/Topology/Exterior.lean index 511879082540a..8b46e6e2b00e1 100644 --- a/Mathlib/Topology/Exterior.lean +++ b/Mathlib/Topology/Exterior.lean @@ -60,21 +60,6 @@ theorem exterior_union (s t : Set X) : exterior (s ∪ t) = exterior s ∪ exter theorem exterior_sUnion (S : Set (Set X)) : exterior (⋃₀ S) = ⋃ s ∈ S, exterior s := by simp only [sUnion_eq_biUnion, exterior_iUnion] -theorem exterior_iInter_subset {s : ι → Set X} : exterior (⋂ i, s i) ⊆ ⋂ i, exterior (s i) := by - simp_rw [exterior] - refine ker_mono (nhdsSet_iInter_le _) |>.trans_eq ?_ - simp_rw [ker_iInf] - -theorem exterior_inter_subset {s t : Set X} : exterior (s ∩ t) ⊆ exterior s ∩ exterior t := by - simp_rw [exterior] - refine ker_mono (nhdsSet_inter_le _ _) |>.trans_eq ?_ - rw [ker_inf _ _] - -theorem exterior_sInter_subset {s : Set (Set X)} : exterior (⋂₀ s) ⊆ ⋂ x ∈ s, exterior x := by - simp_rw [exterior] - refine ker_mono (nhdsSet_sInter_le _) |>.trans_eq ?_ - simp_rw [ker_iInf] - theorem mem_exterior_iff_specializes : x ∈ exterior s ↔ ∃ y ∈ s, x ⤳ y := calc x ∈ exterior s ↔ x ∈ exterior (⋃ y ∈ s, {y}) := by simp _ ↔ ∃ y ∈ s, x ⤳ y := by @@ -98,6 +83,15 @@ theorem exterior_eq_exterior_iff_nhdsSet : exterior s = exterior t ↔ 𝓝ˢ s lemma specializes_iff_exterior_subset : x ⤳ y ↔ exterior {x} ⊆ exterior {y} := by simp [Specializes] +theorem exterior_iInter_subset {s : ι → Set X} : exterior (⋂ i, s i) ⊆ ⋂ i, exterior (s i) := + exterior_mono.map_iInf_le + +theorem exterior_inter_subset {s t : Set X} : exterior (s ∩ t) ⊆ exterior s ∩ exterior t := + exterior_mono.map_inf_le _ _ + +theorem exterior_sInter_subset {s : Set (Set X)} : exterior (⋂₀ s) ⊆ ⋂ x ∈ s, exterior x := + exterior_mono.map_sInf_le + @[simp] lemma exterior_empty : exterior (∅ : Set X) = ∅ := isOpen_empty.exterior_eq @[simp] lemma exterior_univ : exterior (univ : Set X) = univ := isOpen_univ.exterior_eq From a3996466391c60195c5b37cdd1d7496f75962049 Mon Sep 17 00:00:00 2001 From: Yakov Pechersky Date: Sun, 20 Oct 2024 19:39:56 +0000 Subject: [PATCH 208/505] feat(Analysis/Normed/Field/Ultra): nonarchimedean iff norm of nats le one (#15077) Co-authored-by: Yury G. Kudryashov Co-authored-by: Yakov Pechersky Co-authored-by: Yakov Pechersky --- Mathlib.lean | 1 + Mathlib/Analysis/Normed/Field/Ultra.lean | 131 +++++++++++++++++++++++ 2 files changed, 132 insertions(+) create mode 100644 Mathlib/Analysis/Normed/Field/Ultra.lean diff --git a/Mathlib.lean b/Mathlib.lean index ebba90f308a9c..a64206aa580a9 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1239,6 +1239,7 @@ import Mathlib.Analysis.Normed.Field.Basic import Mathlib.Analysis.Normed.Field.InfiniteSum import Mathlib.Analysis.Normed.Field.Lemmas import Mathlib.Analysis.Normed.Field.ProperSpace +import Mathlib.Analysis.Normed.Field.Ultra import Mathlib.Analysis.Normed.Field.UnitBall import Mathlib.Analysis.Normed.Group.AddCircle import Mathlib.Analysis.Normed.Group.AddTorsor diff --git a/Mathlib/Analysis/Normed/Field/Ultra.lean b/Mathlib/Analysis/Normed/Field/Ultra.lean new file mode 100644 index 0000000000000..bbc28a3967fd1 --- /dev/null +++ b/Mathlib/Analysis/Normed/Field/Ultra.lean @@ -0,0 +1,131 @@ +/- +Copyright (c) 2024 Yakov Pechersky. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yakov Pechersky +-/ +import Mathlib.Analysis.Normed.Field.Basic +import Mathlib.Analysis.Normed.Group.Ultra + +/-! +## Sufficient conditions to have an ultrametric norm on a division ring + +This file provides ways of constructing an instance of `IsUltrametricDist` based on +facts about the existing norm. + +## Main results + +* `isUltrametricDist_of_forall_norm_natCast_le_one`: a norm in a division ring is ultrametric + if the norm of the image of a natural is less than or equal to one + +## Implementation details + +The proof relies on a bounded-from-above argument. The main result has a longer proof +to be able to be applied in noncommutative division rings. + +## Tags + +ultrametric, nonarchimedean +-/ +open Metric NNReal + +namespace IsUltrametricDist + +section sufficient + +variable {R : Type*} [NormedDivisionRing R] + +lemma isUltrametricDist_of_forall_norm_add_one_le_max_norm_one + (h : ∀ x : R, ‖x + 1‖ ≤ max ‖x‖ 1) : IsUltrametricDist R := by + refine isUltrametricDist_of_forall_norm_add_le_max_norm (fun x y ↦ ?_) + rcases eq_or_ne y 0 with rfl | hy + · simpa only [add_zero] using le_max_left _ _ + · have p : 0 < ‖y‖ := norm_pos_iff.mpr hy + simpa only [div_add_one hy, norm_div, div_le_iff₀ p, max_mul_of_nonneg _ _ p.le, one_mul, + div_mul_cancel₀ _ p.ne'] using h (x / y) + +lemma isUltrametricDist_of_forall_norm_add_one_of_norm_le_one + (h : ∀ x : R, ‖x‖ ≤ 1 → ‖x + 1‖ ≤ 1) : IsUltrametricDist R := by + refine isUltrametricDist_of_forall_norm_add_one_le_max_norm_one fun x ↦ ?_ + rcases le_or_lt ‖x‖ 1 with H|H + · exact (h _ H).trans (le_max_right _ _) + · suffices ‖x + 1‖ ≤ ‖x‖ from this.trans (le_max_left _ _) + rw [← div_le_div_right (c := ‖x‖) (H.trans' zero_lt_one), div_self (H.trans' zero_lt_one).ne', + ← norm_div, add_div, div_self (by simpa using (H.trans' zero_lt_one)), add_comm] + apply h + simp [inv_le_one_iff₀, H.le] + +lemma isUltrametricDist_of_forall_norm_sub_one_of_norm_le_one + (h : ∀ x : R, ‖x‖ ≤ 1 → ‖x - 1‖ ≤ 1) : IsUltrametricDist R := by + have (x : R) (hx : ‖x‖ ≤ 1) : ‖x + 1‖ ≤ 1 := by + simpa only [← neg_add', norm_neg] using h (-x) (norm_neg x ▸ hx) + exact isUltrametricDist_of_forall_norm_add_one_of_norm_le_one this + +/-- This technical lemma is used in the proof of +`isUltrametricDist_of_forall_norm_natCast_le_one`. -/ +lemma isUltrametricDist_of_forall_pow_norm_le_nsmul_pow_max_one_norm + (h : ∀ (x : R) (m : ℕ), ‖x + 1‖ ^ m ≤ (m + 1) • max 1 (‖x‖ ^ m)) : + IsUltrametricDist R := by + -- it will suffice to prove that `‖x + 1‖ ≤ max 1 ‖x‖` + refine isUltrametricDist_of_forall_norm_add_one_le_max_norm_one fun x ↦ ?_ + -- Morally, we want to deduce this from the hypothesis `h` by taking an `m`-th root and showing + -- that `(m + 1) ^ (1 / m)` gets arbitrarily close to 1, although we will formalise this in a way + -- that avoids explicitly mentioning `m`-th roots. + -- First note it suffices to show that `‖x + 1‖ ≤ a` for all `a : ℝ` with `max ‖x‖ 1 < a`. + rw [max_comm] + refine le_of_forall_le_of_dense fun a ha ↦ ?_ + have ha' : 1 < a := (max_lt_iff.mp ha).left + -- `max 1 ‖x‖ < a`, so there must be some `m : ℕ` such that `m + 1 < (a / max 1 ‖x‖) ^ m` + -- by the virtue of exponential growth being faster than linear growth + obtain ⟨m, hm⟩ : ∃ m : ℕ, ((m + 1) : ℕ) < (a / (max 1 ‖x‖)) ^ m := by + apply_mod_cast Real.exists_natCast_add_one_lt_pow_of_one_lt + rwa [one_lt_div (by positivity)] + -- and we rearrange again to get `(m + 1) • max 1 ‖x‖ ^ m < a ^ m` + rw [div_pow, lt_div_iff₀ (by positivity), ← nsmul_eq_mul] at hm + -- which squeezes down to get our `‖x + 1‖ ≤ a` using our to-be-proven hypothesis of + -- `‖x + 1‖ ^ m ≤ (m + 1) • max 1 ‖x‖ ^ m`, so we're done + -- we can distribute powers into the right term of `max` + have hp : max 1 ‖x‖ ^ m = max 1 (‖x‖ ^ m) := by + have : MonotoneOn (fun a : ℝ ↦ a ^ m) (Set.Ici _) := fun a h b _ h' ↦ pow_le_pow_left h h' _ + rw [this.map_max zero_le_one (norm_nonneg x), one_pow] + rw [hp] at hm + refine le_of_pow_le_pow_left (fun h ↦ ?_) (zero_lt_one.trans ha').le ((h _ _).trans hm.le) + simp only [h, zero_add, pow_zero, max_self, one_smul, lt_self_iff_false] at hm + +/-- To prove that a normed division ring is nonarchimedean, it suffices to prove that the norm +of the image of any natural is less than or equal to one. -/ +lemma isUltrametricDist_of_forall_norm_natCast_le_one + (h : ∀ n : ℕ, ‖(n : R)‖ ≤ 1) : IsUltrametricDist R := by + -- from a previous lemma, suffices to prove that for all `m`, we have + -- `‖x + 1‖ ^ m ≤ (m + 1) • max 1 ‖x‖ ^ m` + refine isUltrametricDist_of_forall_pow_norm_le_nsmul_pow_max_one_norm (fun x m ↦ ?_) + -- we first use our hypothesis about the norm of naturals to have that multiplication by + -- naturals keeps the norm small + replace h (x : R) (n : ℕ) : ‖n • x‖ ≤ ‖x‖ := by + rw [nsmul_eq_mul, norm_mul] + rcases (norm_nonneg x).eq_or_lt with hx | hx + · simp only [← hx, mul_zero, le_refl] + · simpa only [mul_le_iff_le_one_left hx] using h _ + -- we expand the LHS using the binomial theorem, and apply the hypothesis to bound each term by + -- a power of ‖x‖ + transitivity ∑ k ∈ Finset.range (m + 1), ‖x‖ ^ k + · simpa only [← norm_pow, (Commute.one_right x).add_pow, one_pow, mul_one, nsmul_eq_mul, + Nat.cast_comm] using (norm_sum_le _ _).trans (Finset.sum_le_sum fun _ _ ↦ h _ _) + -- the nature of the norm means that one of `1` and `‖x‖ ^ m` is the largest of the two, so the + -- other terms in the binomial expansion are bounded by the max of these, and the number of terms + -- in the sum is precisely `m + 1` + rw [← Finset.card_range (m + 1), ← Finset.sum_const, Finset.card_range] + rcases max_cases 1 (‖x‖ ^ m) with (⟨hm, hx⟩|⟨hm, hx⟩) <;> rw [hm] <;> + -- which we show by comparing the terms in the sum one by one + refine Finset.sum_le_sum fun i hi ↦ ?_ + · rcases eq_or_ne m 0 with rfl | hm + · simp only [pow_zero, le_refl, + show i = 0 by simpa only [zero_add, Finset.range_one, Finset.mem_singleton] using hi] + · rw [pow_le_one_iff_of_nonneg (norm_nonneg _) hm] at hx + exact pow_le_one₀ (norm_nonneg _) hx + · refine pow_le_pow_right₀ ?_ (by simpa only [Finset.mem_range, Nat.lt_succ] using hi) + contrapose! hx + exact pow_le_one₀ (norm_nonneg _) hx.le + +end sufficient + +end IsUltrametricDist From cf0049ec812ebc91c7295af712a53f415308a142 Mon Sep 17 00:00:00 2001 From: Markus Himmel Date: Sun, 20 Oct 2024 19:54:50 +0000 Subject: [PATCH 209/505] feat(CategoryTheory): natural version of `FullyFaithful.homEquiv` (#17450) --- Mathlib/CategoryTheory/Yoneda.lean | 22 +++++++++++++++++++++- 1 file changed, 21 insertions(+), 1 deletion(-) diff --git a/Mathlib/CategoryTheory/Yoneda.lean b/Mathlib/CategoryTheory/Yoneda.lean index 61b37dc22ad22..5116a9b934fd8 100644 --- a/Mathlib/CategoryTheory/Yoneda.lean +++ b/Mathlib/CategoryTheory/Yoneda.lean @@ -24,7 +24,7 @@ namespace CategoryTheory open Opposite -universe v v₁ u₁ u₂ +universe v v₁ v₂ u₁ u₂ -- morphism levels before object levels. See note [CategoryTheory universes]. variable {C : Type u₁} [Category.{v₁} C] @@ -767,4 +767,24 @@ lemma yonedaMap_app_apply {Y : C} {X : Cᵒᵖ} (f : X.unop ⟶ Y) : end +namespace Functor.FullyFaithful + +variable {C : Type u₁} [Category.{v₁} C] + +/-- `FullyFaithful.homEquiv` as a natural isomorphism. -/ +def homNatIso {D : Type u₂} [Category.{v₂} D] {F : C ⥤ D} (hF : F.FullyFaithful) (X : C) : + F.op ⋙ yoneda.obj (F.obj X) ⋙ uliftFunctor.{v₁} ≅ yoneda.obj X ⋙ uliftFunctor.{v₂} := + NatIso.ofComponents + (fun Y => Equiv.toIso (Equiv.ulift.trans <| hF.homEquiv.symm.trans Equiv.ulift.symm)) + (fun f => by ext; exact Equiv.ulift.injective (hF.map_injective (by simp))) + +/-- `FullyFaithful.homEquiv` as a natural isomorphism. -/ +def homNatIsoMaxRight {D : Type u₂} [Category.{max v₁ v₂} D] {F : C ⥤ D} (hF : F.FullyFaithful) + (X : C) : F.op ⋙ yoneda.obj (F.obj X) ≅ yoneda.obj X ⋙ uliftFunctor.{v₂} := + NatIso.ofComponents + (fun Y => Equiv.toIso (hF.homEquiv.symm.trans Equiv.ulift.symm)) + (fun f => by ext; exact Equiv.ulift.injective (hF.map_injective (by simp))) + +end Functor.FullyFaithful + end CategoryTheory From 1b6c758c8ee0b23cbafb1c5add718cd58333c143 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 20 Oct 2024 21:29:44 +0000 Subject: [PATCH 210/505] =?UTF-8?q?feat:=20`negOnePow=20n=20=3D=20(-1=20:?= =?UTF-8?q?=20=E2=84=A4)=20^=20n`=20(#17967)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit and rename `Int.coe_negOnePow` to `Int.cast_negOnePow` for disambiguation --- Mathlib/Algebra/Ring/NegOnePow.lean | 12 +++++++++--- .../SpecialFunctions/Trigonometric/Basic.lean | 12 ++++++------ 2 files changed, 15 insertions(+), 9 deletions(-) diff --git a/Mathlib/Algebra/Ring/NegOnePow.lean b/Mathlib/Algebra/Ring/NegOnePow.lean index 6696881dcbf2e..9f52047d0c32c 100644 --- a/Mathlib/Algebra/Ring/NegOnePow.lean +++ b/Mathlib/Algebra/Ring/NegOnePow.lean @@ -101,11 +101,17 @@ lemma negOnePow_eq_iff (n₁ n₂ : ℤ) : lemma negOnePow_mul_self (n : ℤ) : (n * n).negOnePow = n.negOnePow := by simpa [mul_sub, negOnePow_eq_iff] using n.even_mul_pred_self -lemma coe_negOnePow (K : Type*) (n : ℤ) [Field K] : n.negOnePow = (-1 : K) ^ n := by +lemma cast_negOnePow (K : Type*) (n : ℤ) [Field K] : n.negOnePow = (-1 : K) ^ n := by rcases even_or_odd' n with ⟨k, rfl | rfl⟩ - · rw [zpow_mul, zpow_ofNat] - simp + · simp [zpow_mul, zpow_ofNat] · rw [zpow_add_one₀ (by norm_num), zpow_mul, zpow_ofNat] simp +@[deprecated (since := "2024-10-20")] alias coe_negOnePow := cast_negOnePow + +lemma cast_negOnePow_natCast (R : Type*) [Ring R] (n : ℕ) : negOnePow n = (-1 : R) ^ n := by + obtain ⟨k, rfl | rfl⟩ := Nat.even_or_odd' n <;> simp [pow_succ, pow_mul] + +lemma coe_negOnePow_natCast (n : ℕ) : negOnePow n = (-1 : ℤ) ^ n := cast_negOnePow_natCast .. + end Int diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean index 8e6add75779d2..b09fb6620f62e 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean @@ -280,19 +280,19 @@ theorem sin_int_mul_two_pi_sub (x : ℝ) (n : ℤ) : sin (n * (2 * π) - x) = -s sin_neg x ▸ sin_periodic.int_mul_sub_eq n theorem sin_add_int_mul_pi (x : ℝ) (n : ℤ) : sin (x + n * π) = (-1) ^ n * sin x := - n.coe_negOnePow ℝ ▸ sin_antiperiodic.add_int_mul_eq n + n.cast_negOnePow ℝ ▸ sin_antiperiodic.add_int_mul_eq n theorem sin_add_nat_mul_pi (x : ℝ) (n : ℕ) : sin (x + n * π) = (-1) ^ n * sin x := sin_antiperiodic.add_nat_mul_eq n theorem sin_sub_int_mul_pi (x : ℝ) (n : ℤ) : sin (x - n * π) = (-1) ^ n * sin x := - n.coe_negOnePow ℝ ▸ sin_antiperiodic.sub_int_mul_eq n + n.cast_negOnePow ℝ ▸ sin_antiperiodic.sub_int_mul_eq n theorem sin_sub_nat_mul_pi (x : ℝ) (n : ℕ) : sin (x - n * π) = (-1) ^ n * sin x := sin_antiperiodic.sub_nat_mul_eq n theorem sin_int_mul_pi_sub (x : ℝ) (n : ℤ) : sin (n * π - x) = -((-1) ^ n * sin x) := by - simpa only [sin_neg, mul_neg, Int.coe_negOnePow] using sin_antiperiodic.int_mul_sub_eq n + simpa only [sin_neg, mul_neg, Int.cast_negOnePow] using sin_antiperiodic.int_mul_sub_eq n theorem sin_nat_mul_pi_sub (x : ℝ) (n : ℕ) : sin (n * π - x) = -((-1) ^ n * sin x) := by simpa only [sin_neg, mul_neg] using sin_antiperiodic.nat_mul_sub_eq n @@ -363,19 +363,19 @@ theorem cos_int_mul_two_pi_sub (x : ℝ) (n : ℤ) : cos (n * (2 * π) - x) = co cos_neg x ▸ cos_periodic.int_mul_sub_eq n theorem cos_add_int_mul_pi (x : ℝ) (n : ℤ) : cos (x + n * π) = (-1) ^ n * cos x := - n.coe_negOnePow ℝ ▸ cos_antiperiodic.add_int_mul_eq n + n.cast_negOnePow ℝ ▸ cos_antiperiodic.add_int_mul_eq n theorem cos_add_nat_mul_pi (x : ℝ) (n : ℕ) : cos (x + n * π) = (-1) ^ n * cos x := cos_antiperiodic.add_nat_mul_eq n theorem cos_sub_int_mul_pi (x : ℝ) (n : ℤ) : cos (x - n * π) = (-1) ^ n * cos x := - n.coe_negOnePow ℝ ▸ cos_antiperiodic.sub_int_mul_eq n + n.cast_negOnePow ℝ ▸ cos_antiperiodic.sub_int_mul_eq n theorem cos_sub_nat_mul_pi (x : ℝ) (n : ℕ) : cos (x - n * π) = (-1) ^ n * cos x := cos_antiperiodic.sub_nat_mul_eq n theorem cos_int_mul_pi_sub (x : ℝ) (n : ℤ) : cos (n * π - x) = (-1) ^ n * cos x := - n.coe_negOnePow ℝ ▸ cos_neg x ▸ cos_antiperiodic.int_mul_sub_eq n + n.cast_negOnePow ℝ ▸ cos_neg x ▸ cos_antiperiodic.int_mul_sub_eq n theorem cos_nat_mul_pi_sub (x : ℝ) (n : ℕ) : cos (n * π - x) = (-1) ^ n * cos x := cos_neg x ▸ cos_antiperiodic.nat_mul_sub_eq n From 69bbcb4a9d48465ffcabd3ef9581cc9064df7e06 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 20 Oct 2024 21:50:58 +0000 Subject: [PATCH 211/505] chore: move `mul_ite`, add `dite` version, multiplicativise `ite_add_ite` (#17963) --- Mathlib/Algebra/Group/Defs.lean | 67 +++++++++++++++++++++++++++++++++ Mathlib/Algebra/Ring/Defs.lean | 28 -------------- 2 files changed, 67 insertions(+), 28 deletions(-) diff --git a/Mathlib/Algebra/Group/Defs.lean b/Mathlib/Algebra/Group/Defs.lean index 6a522dad55e3f..119ee7ece01ce 100644 --- a/Mathlib/Algebra/Group/Defs.lean +++ b/Mathlib/Algebra/Group/Defs.lean @@ -168,6 +168,73 @@ class Inv (α : Type u) where @[inherit_doc] postfix:max "⁻¹" => Inv.inv +section ite +variable {α : Type*} (P : Prop) [Decidable P] + +section Mul +variable [Mul α] + +@[to_additive] +lemma mul_dite (a : α) (b : P → α) (c : ¬ P → α) : + (a * if h : P then b h else c h) = if h : P then a * b h else a * c h := by split <;> rfl + +@[to_additive] +lemma mul_ite (a b c : α) : (a * if P then b else c) = if P then a * b else a * c := mul_dite .. + +@[to_additive] +lemma dite_mul (a : P → α) (b : ¬ P → α) (c : α) : + (if h : P then a h else b h) * c = if h : P then a h * c else b h * c := by split <;> rfl + +@[to_additive] +lemma ite_mul (a b c : α) : (if P then a else b) * c = if P then a * c else b * c := dite_mul .. + +-- We make `mul_ite` and `ite_mul` simp lemmas, but not `add_ite` or `ite_add`. +-- The problem we're trying to avoid is dealing with sums of the form `∑ x ∈ s, (f x + ite P 1 0)`, +-- in which `add_ite` followed by `sum_ite` would needlessly slice up +-- the `f x` terms according to whether `P` holds at `x`. +-- There doesn't appear to be a corresponding difficulty so far with `mul_ite` and `ite_mul`. +attribute [simp] mul_dite dite_mul mul_ite ite_mul + +@[to_additive] +lemma dite_mul_dite (a : P → α) (b : ¬ P → α) (c : P → α) (d : ¬ P → α) : + ((if h : P then a h else b h) * if h : P then c h else d h) = + if h : P then a h * c h else b h * d h := by split <;> rfl + +@[to_additive] +lemma ite_mul_ite (a b c d : α) : + ((if P then a else b) * if P then c else d) = if P then a * c else b * d := by split <;> rfl + +end Mul + +section Div +variable [Div α] + +@[to_additive] +lemma div_dite (a : α) (b : P → α) (c : ¬ P → α) : + (a / if h : P then b h else c h) = if h : P then a / b h else a / c h := by split <;> rfl + +@[to_additive] +lemma div_ite (a b c : α) : (a / if P then b else c) = if P then a / b else a / c := div_dite .. + +@[to_additive] +lemma dite_div (a : P → α) (b : ¬ P → α) (c : α) : + (if h : P then a h else b h) / c = if h : P then a h / c else b h / c := by split <;> rfl + +@[to_additive] +lemma ite_div (a b c : α) : (if P then a else b) / c = if P then a / c else b / c := dite_div .. + +@[to_additive] +lemma dite_div_dite (a : P → α) (b : ¬ P → α) (c : P → α) (d : ¬ P → α) : + ((if h : P then a h else b h) / if h : P then c h else d h) = + if h : P then a h / c h else b h / d h := by split <;> rfl + +@[to_additive] +lemma ite_div_ite (a b c d : α) : + ((if P then a else b) / if P then c else d) = if P then a / c else b / d := dite_div_dite .. + +end Div +end ite + section Mul variable [Mul G] diff --git a/Mathlib/Algebra/Ring/Defs.lean b/Mathlib/Algebra/Ring/Defs.lean index 8ff6a13b2f8f9..0fdf58f79d80f 100644 --- a/Mathlib/Algebra/Ring/Defs.lean +++ b/Mathlib/Algebra/Ring/Defs.lean @@ -178,34 +178,6 @@ theorem mul_two (n : α) : n * 2 = n + n := end NonAssocSemiring -@[to_additive] -theorem mul_ite {α} [Mul α] (P : Prop) [Decidable P] (a b c : α) : - (a * if P then b else c) = if P then a * b else a * c := by split_ifs <;> rfl - -@[to_additive] -theorem ite_mul {α} [Mul α] (P : Prop) [Decidable P] (a b c : α) : - (if P then a else b) * c = if P then a * c else b * c := by split_ifs <;> rfl - --- We make `mul_ite` and `ite_mul` simp lemmas, --- but not `add_ite` or `ite_add`. --- The problem we're trying to avoid is dealing with --- summations of the form `∑ x ∈ s, (f x + ite P 1 0)`, --- in which `add_ite` followed by `sum_ite` would needlessly slice up --- the `f x` terms according to whether `P` holds at `x`. --- There doesn't appear to be a corresponding difficulty so far with --- `mul_ite` and `ite_mul`. -attribute [simp] mul_ite ite_mul - -theorem ite_sub_ite {α} [Sub α] (P : Prop) [Decidable P] (a b c d : α) : - ((if P then a else b) - if P then c else d) = if P then a - c else b - d := by - split - repeat rfl - -theorem ite_add_ite {α} [Add α] (P : Prop) [Decidable P] (a b c d : α) : - ((if P then a else b) + if P then c else d) = if P then a + c else b + d := by - split - repeat rfl - section MulZeroClass variable [MulZeroClass α] (P Q : Prop) [Decidable P] [Decidable Q] (a b : α) From 0cfd3c80402e8ef6d0403c4a22cc3d905565fec6 Mon Sep 17 00:00:00 2001 From: damiano Date: Mon, 21 Oct 2024 00:31:07 +0000 Subject: [PATCH 212/505] chore: more fixes for multigoal linter (#17966) Preparation for #12339. --- .../Polynomial/SumIteratedDerivative.lean | 16 ++++++------- .../Morphisms/RingHomProperties.lean | 24 +++++++++---------- Mathlib/Analysis/SpecificLimits/RCLike.lean | 4 ++-- Mathlib/Data/Nat/Choose/Factorization.lean | 4 ++-- Mathlib/NumberTheory/Ostrowski.lean | 3 ++- Mathlib/RingTheory/LocalRing/Quotient.lean | 2 +- Mathlib/RingTheory/Smooth/Pi.lean | 24 +++++++++---------- Mathlib/Topology/MetricSpace/HolderNorm.lean | 2 +- 8 files changed, 40 insertions(+), 39 deletions(-) diff --git a/Mathlib/Algebra/Polynomial/SumIteratedDerivative.lean b/Mathlib/Algebra/Polynomial/SumIteratedDerivative.lean index b8b64c16ec9bf..64723a4d2758b 100644 --- a/Mathlib/Algebra/Polynomial/SumIteratedDerivative.lean +++ b/Mathlib/Algebra/Polynomial/SumIteratedDerivative.lean @@ -198,7 +198,7 @@ theorem aeval_sumIDeriv_of_pos [Nontrivial A] [NoZeroDivisors A] (p : R[X]) {q : rw [map_zero, map_zero, smul_zero, add_zero] rw [Polynomial.map_zero] at hp replace hp := (mul_eq_zero.mp hp.symm).resolve_left ?_ - rw [hp, eval_zero, smul_zero] + · rw [hp, eval_zero, smul_zero] exact fun h => X_sub_C_ne_zero r (pow_eq_zero h) let c k := if hk : q ≤ k then (aeval_iterate_derivative_of_ge A p q hk).choose else 0 have c_le (k) : (c k).natDegree ≤ p.natDegree - k := by @@ -228,13 +228,13 @@ theorem aeval_sumIDeriv_of_pos [Nontrivial A] [NoZeroDivisors A] (p : R[X]) {q : rw [sumIDeriv_apply, map_sum, map_sum, this] have : range q = range (q - 1 + 1) := by rw [tsub_add_cancel_of_le (Nat.one_le_of_lt hq)] rw [sum_union, this, sum_range_succ] - congr 2 - · apply sum_eq_zero - exact fun x hx => aeval_iterate_derivative_of_lt p _ r hp (mem_range.mp hx) - · rw [← aeval_iterate_derivative_self _ _ _ hp] - · rw [smul_sum, sum_congr rfl] - intro k hk - exact hc k (mem_Ico.mp hk).1 r + · congr 2 + · apply sum_eq_zero + exact fun x hx => aeval_iterate_derivative_of_lt p _ r hp (mem_range.mp hx) + · rw [← aeval_iterate_derivative_self _ _ _ hp] + · rw [smul_sum, sum_congr rfl] + intro k hk + exact hc k (mem_Ico.mp hk).1 r · rw [range_eq_Ico, disjoint_iff_inter_eq_empty, eq_empty_iff_forall_not_mem] intro x hx rw [mem_inter, mem_Ico, mem_Ico] at hx diff --git a/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean b/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean index 06bd0001fe507..06f83cb90fe8a 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean @@ -444,18 +444,18 @@ private lemma respects_isOpenImmersion_aux {X Y : Scheme.{u}} [IsAffine Y] {U : · obtain ⟨(Us : Set Y.Opens), hUs, heq⟩ := Opens.isBasis_iff_cover.mp (isBasis_basicOpen Y) U let V (s : Us) : X.Opens := f ⁻¹ᵁ U.ι ⁻¹ᵁ s rw [IsLocalAtSource.iff_of_iSup_eq_top (P := P) V] - intro s - let f' : (V s).toScheme ⟶ U.ι ⁻¹ᵁ s := f ∣_ U.ι ⁻¹ᵁ s - have hf' : P f' := IsLocalAtTarget.restrict hf _ - let e : (U.ι ⁻¹ᵁ s).toScheme ≅ s := IsOpenImmersion.isoOfRangeEq ((U.ι ⁻¹ᵁ s).ι ≫ U.ι) s.1.ι - (by simpa [Set.range_comp, Set.image_preimage_eq_iff, heq] using le_sSup s.2) - have heq : (V s).ι ≫ f ≫ U.ι = f' ≫ e.hom ≫ s.1.ι := by - simp only [IsOpenImmersion.isoOfRangeEq_hom_fac, f', e, morphismRestrict_ι_assoc] - rw [heq, ← Category.assoc] - refine this _ ?_ ?_ - · rwa [P.cancel_right_of_respectsIso] - · obtain ⟨a, ha⟩ := hUs s.2 - use a, ha.symm + · intro s + let f' : (V s).toScheme ⟶ U.ι ⁻¹ᵁ s := f ∣_ U.ι ⁻¹ᵁ s + have hf' : P f' := IsLocalAtTarget.restrict hf _ + let e : (U.ι ⁻¹ᵁ s).toScheme ≅ s := IsOpenImmersion.isoOfRangeEq ((U.ι ⁻¹ᵁ s).ι ≫ U.ι) s.1.ι + (by simpa [Set.range_comp, Set.image_preimage_eq_iff, heq] using le_sSup s.2) + have heq : (V s).ι ≫ f ≫ U.ι = f' ≫ e.hom ≫ s.1.ι := by + simp only [IsOpenImmersion.isoOfRangeEq_hom_fac, f', e, morphismRestrict_ι_assoc] + rw [heq, ← Category.assoc] + refine this _ ?_ ?_ + · rwa [P.cancel_right_of_respectsIso] + · obtain ⟨a, ha⟩ := hUs s.2 + use a, ha.symm · apply f.preimage_iSup_eq_top apply U.ι.image_injective simp only [U.ι.image_iSup, U.ι.image_preimage_eq_opensRange_inter, Scheme.Opens.opensRange_ι] diff --git a/Mathlib/Analysis/SpecificLimits/RCLike.lean b/Mathlib/Analysis/SpecificLimits/RCLike.lean index 2f82b7954f882..0207c4041ec9f 100644 --- a/Mathlib/Analysis/SpecificLimits/RCLike.lean +++ b/Mathlib/Analysis/SpecificLimits/RCLike.lean @@ -29,8 +29,8 @@ theorem RCLike.tendsto_add_mul_div_add_mul_atTop_nhds (a b c : 𝕜) {d : 𝕜} Tendsto (fun k : ℕ ↦ (a + c * k) / (b + d * k)) atTop (𝓝 (c / d)) := by apply Filter.Tendsto.congr' case f₁ => exact fun k ↦ (a * (↑k)⁻¹ + c) / (b * (↑k)⁻¹ + d) - refine (eventually_ne_atTop 0).mp (Eventually.of_forall ?_) - · intro h hx + · refine (eventually_ne_atTop 0).mp (Eventually.of_forall ?_) + intro h hx field_simp [hx] · apply Filter.Tendsto.div _ _ hd all_goals diff --git a/Mathlib/Data/Nat/Choose/Factorization.lean b/Mathlib/Data/Nat/Choose/Factorization.lean index 0c80076046db8..6bd3a1dac00f3 100644 --- a/Mathlib/Data/Nat/Choose/Factorization.lean +++ b/Mathlib/Data/Nat/Choose/Factorization.lean @@ -40,8 +40,8 @@ theorem factorization_choose_le_log : (choose n k).factorization p ≤ log p n : simp [choose_eq_zero_of_lt hnk] rw [factorization_def _ hp, @padicValNat_def _ ⟨hp⟩ _ (choose_pos hkn)] rw [← Nat.cast_le (α := ℕ∞), ← multiplicity.Finite.emultiplicity_eq_multiplicity] - simp only [hp.emultiplicity_choose hkn (lt_add_one _), Nat.cast_le] - exact (Finset.card_filter_le _ _).trans (le_of_eq (Nat.card_Ico _ _)) + · simp only [hp.emultiplicity_choose hkn (lt_add_one _), Nat.cast_le] + exact (Finset.card_filter_le _ _).trans (le_of_eq (Nat.card_Ico _ _)) apply Nat.multiplicity_finite_iff.2 ⟨hp.ne_one, choose_pos hkn⟩ /-- A `pow` form of `Nat.factorization_choose_le` -/ diff --git a/Mathlib/NumberTheory/Ostrowski.lean b/Mathlib/NumberTheory/Ostrowski.lean index cd549970b163c..03a573e8490f5 100644 --- a/Mathlib/NumberTheory/Ostrowski.lean +++ b/Mathlib/NumberTheory/Ostrowski.lean @@ -335,7 +335,8 @@ lemma one_lt_of_not_bounded (notbdd : ¬ ∀ n : ℕ, f n ≤ 1) {n₀ : ℕ} (h rw [← mul_rpow (by positivity) (by positivity), mul_assoc, add_mul, one_mul, mul_comm _ (k : ℝ)] -- For 0 < logb n₀ n below we also need to exclude n = 1. - rcases eq_or_ne n 1 with rfl | h₁; simp only [Nat.cast_one, map_one, le_refl] + rcases eq_or_ne n 1 with rfl | h₁ + · simp only [Nat.cast_one, map_one, le_refl] refine le_of_tendsto_of_tendsto tendsto_const_nhds ?_ (eventually_atTop.2 ⟨1, h_ineq2⟩) nth_rw 2 [← mul_one 1] have : 0 < logb n₀ n := logb_pos (mod_cast hn₀) (by norm_cast; omega) diff --git a/Mathlib/RingTheory/LocalRing/Quotient.lean b/Mathlib/RingTheory/LocalRing/Quotient.lean index e4ef8ab01d8e7..cc02b61745b60 100644 --- a/Mathlib/RingTheory/LocalRing/Quotient.lean +++ b/Mathlib/RingTheory/LocalRing/Quotient.lean @@ -67,7 +67,7 @@ theorem finrank_quotient_map : · let b := Module.Free.chooseBasis (R ⧸ p) (S ⧸ pS) choose b' hb' using fun i ↦ Ideal.Quotient.mk_surjective (b i) conv_rhs => rw [finrank_eq_card_chooseBasisIndex] - apply finrank_le_of_span_eq_top + refine finrank_le_of_span_eq_top (v := b') ?_ apply (quotient_span_eq_top_iff_span_eq_top _).mp rw [← Set.range_comp, show Ideal.Quotient.mk pS ∘ b' = ⇑b from funext hb'] exact b.span_eq diff --git a/Mathlib/RingTheory/Smooth/Pi.lean b/Mathlib/RingTheory/Smooth/Pi.lean index acebf93edbce4..cfb4cf1dd16df 100644 --- a/Mathlib/RingTheory/Smooth/Pi.lean +++ b/Mathlib/RingTheory/Smooth/Pi.lean @@ -28,18 +28,18 @@ theorem of_pi [FormallySmooth R (Π i, A i)] (i) : FormallySmooth R (A i) := by classical fapply FormallySmooth.of_split (Pi.evalAlgHom R A i) - apply AlgHom.ofLinearMap - ((Ideal.Quotient.mkₐ R _).toLinearMap.comp (LinearMap.single _ _ i)) - · show Ideal.Quotient.mk _ (Pi.single i 1) = 1 - rw [← (Ideal.Quotient.mk _).map_one, ← sub_eq_zero, ← map_sub, - Ideal.Quotient.eq_zero_iff_mem] - have : Pi.single i 1 - 1 ∈ RingHom.ker (Pi.evalAlgHom R A i).toRingHom := by - simp [RingHom.mem_ker] - convert neg_mem (Ideal.pow_mem_pow this 2) using 1 - simp [pow_two, sub_mul, mul_sub, ← Pi.single_mul] - · intro x y - show Ideal.Quotient.mk _ _ = Ideal.Quotient.mk _ _ * Ideal.Quotient.mk _ _ - simp only [AlgHom.toRingHom_eq_coe, LinearMap.coe_single, Pi.single_mul, map_mul] + · apply AlgHom.ofLinearMap + ((Ideal.Quotient.mkₐ R _).toLinearMap.comp (LinearMap.single _ _ i)) + · show Ideal.Quotient.mk _ (Pi.single i 1) = 1 + rw [← (Ideal.Quotient.mk _).map_one, ← sub_eq_zero, ← map_sub, + Ideal.Quotient.eq_zero_iff_mem] + have : Pi.single i 1 - 1 ∈ RingHom.ker (Pi.evalAlgHom R A i).toRingHom := by + simp [RingHom.mem_ker] + convert neg_mem (Ideal.pow_mem_pow this 2) using 1 + simp [pow_two, sub_mul, mul_sub, ← Pi.single_mul] + · intro x y + show Ideal.Quotient.mk _ _ = Ideal.Quotient.mk _ _ * Ideal.Quotient.mk _ _ + simp only [AlgHom.toRingHom_eq_coe, LinearMap.coe_single, Pi.single_mul, map_mul] · ext x show (Pi.single i x) i = x simp diff --git a/Mathlib/Topology/MetricSpace/HolderNorm.lean b/Mathlib/Topology/MetricSpace/HolderNorm.lean index 864d59f082ad4..424128d4adb85 100644 --- a/Mathlib/Topology/MetricSpace/HolderNorm.lean +++ b/Mathlib/Topology/MetricSpace/HolderNorm.lean @@ -158,7 +158,7 @@ lemma MemHolder.holderWith {r : ℝ≥0} {f : X → Y} (hf : MemHolder r f) : by_cases hx : x₁ = x₂ · simp only [hx, edist_self, zero_le] rw [nnHolderNorm, eHolderNorm, coe_toNNReal] - swap; exact hf.eHolderNorm_lt_top.ne + on_goal 2 => exact hf.eHolderNorm_lt_top.ne have h₁ : edist x₁ x₂ ^ (r : ℝ) ≠ 0 := (Ne.symm <| ne_of_lt <| ENNReal.rpow_pos (edist_pos.2 hx) (edist_lt_top x₁ x₂).ne) have h₂ : edist x₁ x₂ ^ (r : ℝ) ≠ ∞ := by From a9827c1f868e43dfe152895f3a0a7b2849f7d44a Mon Sep 17 00:00:00 2001 From: Pietro Monticone <38562595+pitmonticone@users.noreply.github.com> Date: Mon, 21 Oct 2024 00:40:39 +0000 Subject: [PATCH 213/505] feat(Data/Set/Basic): add `diff_insert_of_not_mem` (#17973) Co-authored-by: @Command-Master Upstreamed from the [EquationalTheories](https://github.com/teorth/equational_theories) project. --- Mathlib/Data/Set/Basic.lean | 5 +++++ 1 file changed, 5 insertions(+) diff --git a/Mathlib/Data/Set/Basic.lean b/Mathlib/Data/Set/Basic.lean index 63e22e76e74a8..6c5861316434a 100644 --- a/Mathlib/Data/Set/Basic.lean +++ b/Mathlib/Data/Set/Basic.lean @@ -1561,6 +1561,11 @@ theorem diff_compl : s \ tᶜ = s ∩ t := theorem diff_diff_right {s t u : Set α} : s \ (t \ u) = s \ t ∪ s ∩ u := sdiff_sdiff_right' +theorem diff_insert_of_not_mem {x : α} (h : x ∉ s) : s \ insert x t = s \ t := by + refine Subset.antisymm (diff_subset_diff (refl _) (subset_insert ..)) fun y hy ↦ ?_ + simp only [mem_diff, mem_insert_iff, not_or] at hy ⊢ + exact ⟨hy.1, fun hxy ↦ h <| hxy ▸ hy.1, hy.2⟩ + @[simp] theorem insert_diff_of_mem (s) (h : a ∈ t) : insert a s \ t = s \ t := by ext From 4f5e6c178cc931a23492c3f57699aba5e89dacc5 Mon Sep 17 00:00:00 2001 From: Pietro Monticone <38562595+pitmonticone@users.noreply.github.com> Date: Mon, 21 Oct 2024 01:17:40 +0000 Subject: [PATCH 214/505] feat(Data/Nat/Bitwise): add simp lemmas (#17977) Co-authored-by: @Command-Master Upstreamed from the [EquationalTheories](https://github.com/teorth/equational_theories) project. --- Mathlib/Data/Nat/Bitwise.lean | 16 +++++++++++++++- 1 file changed, 15 insertions(+), 1 deletion(-) diff --git a/Mathlib/Data/Nat/Bitwise.lean b/Mathlib/Data/Nat/Bitwise.lean index c0cbec11309d7..3c394416fac8e 100644 --- a/Mathlib/Data/Nat/Bitwise.lean +++ b/Mathlib/Data/Nat/Bitwise.lean @@ -37,7 +37,6 @@ should be connected. bitwise, and, or, xor -/ - open Function namespace Nat @@ -354,6 +353,21 @@ theorem lt_xor_cases {a b c : ℕ} (h : a < b ^^^ c) : a ^^^ c < b ∨ a ^^^ b < obtain ha | hb | hc := xor_trichotomy <| Nat.xor_assoc _ _ _ ▸ xor_ne_zero.2 h.ne exacts [(h.asymm ha).elim, Or.inl <| Nat.xor_comm _ _ ▸ hb, Or.inr hc] +@[simp] +theorem xor_mod_two_eq {m n : ℕ} : (m ^^^ n) % 2 = (m + n) % 2 := by + by_cases h : (m + n) % 2 = 0 + · simp only [h, mod_two_eq_zero_iff_testBit_zero, testBit_zero, xor_mod_two_eq_one, decide_not, + Bool.decide_iff_dist, Bool.not_eq_false', beq_iff_eq, decide_eq_decide] + omega + · simp only [mod_two_ne_zero] at h + simp only [h, xor_mod_two_eq_one] + omega + +@[simp] +theorem even_xor {m n : ℕ} : Even (m ^^^ n) ↔ (Even m ↔ Even n) := by + simp only [even_iff, xor_mod_two_eq] + omega + @[simp] theorem bit_lt_two_pow_succ_iff {b x n} : bit b x < 2 ^ (n + 1) ↔ x < 2 ^ n := by cases b <;> simp <;> omega From 34309b691800dfc483f72fd46aa73fe1e68c4c5e Mon Sep 17 00:00:00 2001 From: Pietro Monticone <38562595+pitmonticone@users.noreply.github.com> Date: Mon, 21 Oct 2024 01:17:41 +0000 Subject: [PATCH 215/505] chore(build.in): update version of `gh-get-current-pr` action (#17978) This PR updates the version of the `gh-get-current-pr` action from `v2.2.0` to `3.0.0` to avoid the following deprecation warning: > The following actions use a deprecated Node.js version and will be forced to run on node20: 8BitJonny/gh-get-current-pr@2.2.0. For more info: https://github.blog/changelog/2024-03-07-github-actions-all-actions-will-run-on-node20-instead-of-node16-by-default/ --- .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 3facc8b60a92b..3b635701d4469 100644 --- a/.github/build.in.yml +++ b/.github/build.in.yml @@ -303,7 +303,7 @@ jobs: - uses: actions/checkout@v4 - id: PR - uses: 8BitJonny/gh-get-current-pr@2.2.0 + uses: 8BitJonny/gh-get-current-pr@3.0.0 # TODO: this may not work properly if the same commit is pushed to multiple branches: # https://github.com/8BitJonny/gh-get-current-pr/issues/8 with: diff --git a/.github/workflows/bors.yml b/.github/workflows/bors.yml index 207aee8482616..4b91df5b00120 100644 --- a/.github/workflows/bors.yml +++ b/.github/workflows/bors.yml @@ -313,7 +313,7 @@ jobs: - uses: actions/checkout@v4 - id: PR - uses: 8BitJonny/gh-get-current-pr@2.2.0 + uses: 8BitJonny/gh-get-current-pr@3.0.0 # TODO: this may not work properly if the same commit is pushed to multiple branches: # https://github.com/8BitJonny/gh-get-current-pr/issues/8 with: diff --git a/.github/workflows/build.yml b/.github/workflows/build.yml index c47cd306ed4da..d1e5048846e07 100644 --- a/.github/workflows/build.yml +++ b/.github/workflows/build.yml @@ -320,7 +320,7 @@ jobs: - uses: actions/checkout@v4 - id: PR - uses: 8BitJonny/gh-get-current-pr@2.2.0 + uses: 8BitJonny/gh-get-current-pr@3.0.0 # TODO: this may not work properly if the same commit is pushed to multiple branches: # https://github.com/8BitJonny/gh-get-current-pr/issues/8 with: diff --git a/.github/workflows/build_fork.yml b/.github/workflows/build_fork.yml index c431ad5e7d2f2..e3a3e5d64931d 100644 --- a/.github/workflows/build_fork.yml +++ b/.github/workflows/build_fork.yml @@ -317,7 +317,7 @@ jobs: - uses: actions/checkout@v4 - id: PR - uses: 8BitJonny/gh-get-current-pr@2.2.0 + uses: 8BitJonny/gh-get-current-pr@3.0.0 # TODO: this may not work properly if the same commit is pushed to multiple branches: # https://github.com/8BitJonny/gh-get-current-pr/issues/8 with: From 4d87f692a546837be1bd8d1dcd18a36e4722c5bc Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 21 Oct 2024 03:20:56 +0000 Subject: [PATCH 216/505] chore: rename List.modifyNth->modify and insertNth->insertIdx (#17985) This relies on a branch of Batteries, so needs a delegation not a merge. See zulip [discussion](https://leanprover.zulipchat.com/#narrow/channel/348111-batteries/topic/.60List.2EmodifyNth.60/near/477340629). --- Mathlib.lean | 1 + Mathlib/Data/List/Basic.lean | 40 +++-- Mathlib/Data/List/InsertIdx.lean | 203 ++++++++++++++++++++++ Mathlib/Data/List/InsertNth.lean | 188 +------------------- Mathlib/Data/List/Lemmas.lean | 11 +- Mathlib/Data/List/Permutation.lean | 38 ++-- Mathlib/Data/Vector/Basic.lean | 60 ++++--- Mathlib/SetTheory/Game/PGame.lean | 2 +- Mathlib/Testing/SlimCheck/Gen.lean | 2 +- Mathlib/Testing/SlimCheck/Sampleable.lean | 2 +- Mathlib/Topology/List.lean | 54 +++--- lake-manifest.json | 4 +- lakefile.lean | 2 +- scripts/noshake.json | 2 +- test/rewrites.lean | 2 +- 15 files changed, 338 insertions(+), 273 deletions(-) create mode 100644 Mathlib/Data/List/InsertIdx.lean diff --git a/Mathlib.lean b/Mathlib.lean index a64206aa580a9..e722c84da9b0e 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2347,6 +2347,7 @@ import Mathlib.Data.List.GetD import Mathlib.Data.List.GroupBy import Mathlib.Data.List.Indexes import Mathlib.Data.List.Infix +import Mathlib.Data.List.InsertIdx import Mathlib.Data.List.InsertNth import Mathlib.Data.List.Intervals import Mathlib.Data.List.Iterate diff --git a/Mathlib/Data/List/Basic.lean b/Mathlib/Data/List/Basic.lean index 2d8dc25390035..3f6d1c26e9f52 100644 --- a/Mathlib/Data/List/Basic.lean +++ b/Mathlib/Data/List/Basic.lean @@ -837,33 +837,43 @@ theorem eq_cons_of_length_one {l : List α} (h : l.length = 1) : l = [l.get ⟨0 end deprecated -theorem modifyNthTail_modifyNthTail {f g : List α → List α} (m : ℕ) : +theorem modifyTailIdx_modifyTailIdx {f g : List α → List α} (m : ℕ) : ∀ (n) (l : List α), - (l.modifyNthTail f n).modifyNthTail g (m + n) = - l.modifyNthTail (fun l => (f l).modifyNthTail g m) n + (l.modifyTailIdx f n).modifyTailIdx g (m + n) = + l.modifyTailIdx (fun l => (f l).modifyTailIdx g m) n | 0, _ => rfl | _ + 1, [] => rfl - | n + 1, a :: l => congr_arg (List.cons a) (modifyNthTail_modifyNthTail m n l) + | n + 1, a :: l => congr_arg (List.cons a) (modifyTailIdx_modifyTailIdx m n l) -theorem modifyNthTail_modifyNthTail_le {f g : List α → List α} (m n : ℕ) (l : List α) +@[deprecated (since := "2024-10-21")] +alias modifyNthTail_modifyNthTail := modifyTailIdx_modifyTailIdx + +theorem modifyTailIdx_modifyTailIdx_le {f g : List α → List α} (m n : ℕ) (l : List α) (h : n ≤ m) : - (l.modifyNthTail f n).modifyNthTail g m = - l.modifyNthTail (fun l => (f l).modifyNthTail g (m - n)) n := by + (l.modifyTailIdx f n).modifyTailIdx g m = + l.modifyTailIdx (fun l => (f l).modifyTailIdx g (m - n)) n := by rcases Nat.exists_eq_add_of_le h with ⟨m, rfl⟩ - rw [Nat.add_comm, modifyNthTail_modifyNthTail, Nat.add_sub_cancel] + rw [Nat.add_comm, modifyTailIdx_modifyTailIdx, Nat.add_sub_cancel] + +@[deprecated (since := "2024-10-21")] +alias modifyNthTail_modifyNthTail_le := modifyTailIdx_modifyTailIdx_le -theorem modifyNthTail_modifyNthTail_same {f g : List α → List α} (n : ℕ) (l : List α) : - (l.modifyNthTail f n).modifyNthTail g n = l.modifyNthTail (g ∘ f) n := by - rw [modifyNthTail_modifyNthTail_le n n l (le_refl n), Nat.sub_self]; rfl +theorem modifyTailIdx_modifyTailIdx_same {f g : List α → List α} (n : ℕ) (l : List α) : + (l.modifyTailIdx f n).modifyTailIdx g n = l.modifyTailIdx (g ∘ f) n := by + rw [modifyTailIdx_modifyTailIdx_le n n l (le_refl n), Nat.sub_self]; rfl -@[deprecated (since := "2024-05-04")] alias removeNth_eq_nthTail := eraseIdx_eq_modifyNthTail +@[deprecated (since := "2024-10-21")] +alias modifyNthTail_modifyNthTail_same := modifyTailIdx_modifyTailIdx_same +@[deprecated (since := "2024-05-04")] alias removeNth_eq_nthTail := eraseIdx_eq_modifyTailIdx -theorem modifyNth_eq_set (f : α → α) : - ∀ (n) (l : List α), modifyNth f n l = ((fun a => set l n (f a)) <$> l[n]?).getD l +theorem modify_eq_set (f : α → α) : + ∀ (n) (l : List α), modify f n l = ((fun a => set l n (f a)) <$> l[n]?).getD l | 0, l => by cases l <;> simp | _ + 1, [] => rfl | n + 1, b :: l => - (congr_arg (cons b) (modifyNth_eq_set f n l)).trans <| by cases h : l[n]? <;> simp [h] + (congr_arg (cons b) (modify_eq_set f n l)).trans <| by cases h : l[n]? <;> simp [h] + +@[deprecated (since := "2024-10-21")] alias modifyNth_eq_set := modify_eq_set @[simp] theorem getElem_set_of_ne {l : List α} {i j : ℕ} (h : i ≠ j) (a : α) diff --git a/Mathlib/Data/List/InsertIdx.lean b/Mathlib/Data/List/InsertIdx.lean new file mode 100644 index 0000000000000..7c1c61e68cdc2 --- /dev/null +++ b/Mathlib/Data/List/InsertIdx.lean @@ -0,0 +1,203 @@ +/- +Copyright (c) 2014 Parikshit Khanna. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Parikshit Khanna, Jeremy Avigad, Leonardo de Moura, Floris van Doorn, Mario Carneiro +-/ +import Mathlib.Data.List.Basic + +/-! +# insertIdx + +Proves various lemmas about `List.insertIdx`. +-/ + +assert_not_exists Set.range + +open Function + +open Nat hiding one_pos + +namespace List + +universe u + +variable {α : Type u} + +section InsertIdx + +variable {a : α} + +@[simp] +theorem insertIdx_zero (s : List α) (x : α) : insertIdx 0 x s = x :: s := + rfl + +@[simp] +theorem insertIdx_succ_nil (n : ℕ) (a : α) : insertIdx (n + 1) a [] = [] := + rfl + +@[simp] +theorem insertIdx_succ_cons (s : List α) (hd x : α) (n : ℕ) : + insertIdx (n + 1) x (hd :: s) = hd :: insertIdx n x s := + rfl + +theorem length_insertIdx : ∀ n as, n ≤ length as → length (insertIdx n a as) = length as + 1 + | 0, _, _ => rfl + | _ + 1, [], h => (Nat.not_succ_le_zero _ h).elim + | n + 1, _ :: as, h => congr_arg Nat.succ <| length_insertIdx n as (Nat.le_of_succ_le_succ h) + +theorem eraseIdx_insertIdx (n : ℕ) (l : List α) : (l.insertIdx n a).eraseIdx n = l := by + rw [eraseIdx_eq_modifyTailIdx, insertIdx, modifyTailIdx_modifyTailIdx_same] + exact modifyTailIdx_id _ _ + +theorem insertIdx_eraseIdx_of_ge : + ∀ n m as, + n < length as → n ≤ m → insertIdx m a (as.eraseIdx n) = (as.insertIdx (m + 1) a).eraseIdx n + | 0, 0, [], has, _ => (lt_irrefl _ has).elim + | 0, 0, _ :: as, _, _ => by simp [eraseIdx, insertIdx] + | 0, _ + 1, _ :: _, _, _ => rfl + | n + 1, m + 1, a :: as, has, hmn => + congr_arg (cons a) <| + insertIdx_eraseIdx_of_ge n m as (Nat.lt_of_succ_lt_succ has) (Nat.le_of_succ_le_succ hmn) + +theorem insertIdx_eraseIdx_of_le : + ∀ n m as, + n < length as → m ≤ n → insertIdx m a (as.eraseIdx n) = (as.insertIdx m a).eraseIdx (n + 1) + | _, 0, _ :: _, _, _ => rfl + | n + 1, m + 1, a :: as, has, hmn => + congr_arg (cons a) <| + insertIdx_eraseIdx_of_le n m as (Nat.lt_of_succ_lt_succ has) (Nat.le_of_succ_le_succ hmn) + +theorem insertIdx_comm (a b : α) : + ∀ (i j : ℕ) (l : List α) (_ : i ≤ j) (_ : j ≤ length l), + (l.insertIdx i a).insertIdx (j + 1) b = (l.insertIdx j b).insertIdx i a + | 0, j, l => by simp [insertIdx] + | _ + 1, 0, _ => fun h => (Nat.not_lt_zero _ h).elim + | i + 1, j + 1, [] => by simp + | i + 1, j + 1, c :: l => fun h₀ h₁ => by + simp only [insertIdx_succ_cons, cons.injEq, true_and] + exact insertIdx_comm a b i j l (Nat.le_of_succ_le_succ h₀) (Nat.le_of_succ_le_succ h₁) + +theorem mem_insertIdx {a b : α} : + ∀ {n : ℕ} {l : List α} (_ : n ≤ l.length), a ∈ l.insertIdx n b ↔ a = b ∨ a ∈ l + | 0, as, _ => by simp + | _ + 1, [], h => (Nat.not_succ_le_zero _ h).elim + | n + 1, a' :: as, h => by + rw [List.insertIdx_succ_cons, mem_cons, mem_insertIdx (Nat.le_of_succ_le_succ h), + ← or_assoc, @or_comm (a = a'), or_assoc, mem_cons] + +theorem insertIdx_of_length_lt (l : List α) (x : α) (n : ℕ) (h : l.length < n) : + insertIdx n x l = l := by + induction' l with hd tl IH generalizing n + · cases n + · simp at h + · simp + · cases n + · simp at h + · simp only [Nat.succ_lt_succ_iff, length] at h + simpa using IH _ h + +@[simp] +theorem insertIdx_length_self (l : List α) (x : α) : insertIdx l.length x l = l ++ [x] := by + induction' l with hd tl IH + · simp + · simpa using IH + +theorem length_le_length_insertIdx (l : List α) (x : α) (n : ℕ) : + l.length ≤ (insertIdx n x l).length := by + rcases le_or_lt n l.length with hn | hn + · rw [length_insertIdx _ _ hn] + exact (Nat.lt_succ_self _).le + · rw [insertIdx_of_length_lt _ _ _ hn] + +theorem length_insertIdx_le_succ (l : List α) (x : α) (n : ℕ) : + (insertIdx n x l).length ≤ l.length + 1 := by + rcases le_or_lt n l.length with hn | hn + · rw [length_insertIdx _ _ hn] + · rw [insertIdx_of_length_lt _ _ _ hn] + exact (Nat.lt_succ_self _).le + +theorem getElem_insertIdx_of_lt (l : List α) (x : α) (n k : ℕ) (hn : k < n) (hk : k < l.length) + (hk' : k < (insertIdx n x l).length := hk.trans_le (length_le_length_insertIdx _ _ _)) : + (insertIdx n x l)[k] = l[k] := by + induction' n with n IH generalizing k l + · simp at hn + · cases' l with hd tl + · simp + · cases k + · simp [get] + · rw [Nat.succ_lt_succ_iff] at hn + simpa using IH _ _ hn _ + +theorem get_insertIdx_of_lt (l : List α) (x : α) (n k : ℕ) (hn : k < n) (hk : k < l.length) + (hk' : k < (insertIdx n x l).length := hk.trans_le (length_le_length_insertIdx _ _ _)) : + (insertIdx n x l).get ⟨k, hk'⟩ = l.get ⟨k, hk⟩ := by + simp_all [getElem_insertIdx_of_lt] + +@[simp] +theorem getElem_insertIdx_self (l : List α) (x : α) (n : ℕ) (hn : n ≤ l.length) + (hn' : n < (insertIdx n x l).length := (by rwa [length_insertIdx _ _ hn, Nat.lt_succ_iff])) : + (insertIdx n x l)[n] = x := by + induction' l with hd tl IH generalizing n + · simp only [length] at hn + cases hn + simp only [insertIdx_zero, getElem_singleton] + · cases n + · simp + · simp only [Nat.succ_le_succ_iff, length] at hn + simpa using IH _ hn + +theorem get_insertIdx_self (l : List α) (x : α) (n : ℕ) (hn : n ≤ l.length) + (hn' : n < (insertIdx n x l).length := (by rwa [length_insertIdx _ _ hn, Nat.lt_succ_iff])) : + (insertIdx n x l).get ⟨n, hn'⟩ = x := by + simp [hn, hn'] + +theorem getElem_insertIdx_add_succ (l : List α) (x : α) (n k : ℕ) (hk' : n + k < l.length) + (hk : n + k + 1 < (insertIdx n x l).length := (by + rwa [length_insertIdx _ _ (by omega), Nat.succ_lt_succ_iff])) : + (insertIdx n x l)[n + k + 1] = l[n + k] := by + induction' l with hd tl IH generalizing n k + · simp at hk' + · cases n + · simp + · simpa [Nat.add_right_comm] using IH _ _ _ + +theorem get_insertIdx_add_succ (l : List α) (x : α) (n k : ℕ) (hk' : n + k < l.length) + (hk : n + k + 1 < (insertIdx n x l).length := (by + rwa [length_insertIdx _ _ (by omega), Nat.succ_lt_succ_iff])) : + (insertIdx n x l).get ⟨n + k + 1, hk⟩ = get l ⟨n + k, hk'⟩ := by + simp [getElem_insertIdx_add_succ, hk, hk'] + +set_option linter.unnecessarySimpa false in +theorem insertIdx_injective (n : ℕ) (x : α) : Function.Injective (insertIdx n x) := by + induction' n with n IH + · have : insertIdx 0 x = cons x := funext fun _ => rfl + simp [this] + · rintro (_ | ⟨a, as⟩) (_ | ⟨b, bs⟩) h <;> simpa [IH.eq_iff] using h + +@[deprecated (since := "2024-10-21")] alias insertNth_zero := insertIdx_zero +@[deprecated (since := "2024-10-21")] alias insertNth_succ_nil := insertIdx_succ_nil +@[deprecated (since := "2024-10-21")] alias insertNth_succ_cons := insertIdx_succ_cons +@[deprecated (since := "2024-10-21")] alias length_insertNth := length_insertIdx +@[deprecated (since := "2024-10-21")] alias removeNth_insertIdx := eraseIdx_insertIdx +@[deprecated (since := "2024-05-04")] alias removeNth_insertNth := eraseIdx_insertIdx +@[deprecated (since := "2024-10-21")] alias insertNth_eraseIdx_of_ge := insertIdx_eraseIdx_of_ge +@[deprecated (since := "2024-05-04")] alias insertNth_removeNth_of_ge := insertIdx_eraseIdx_of_ge +@[deprecated (since := "2024-10-21")] alias insertNth_eraseIdx_of_le := insertIdx_eraseIdx_of_le +@[deprecated (since := "2024-05-04")] alias insertIdx_removeNth_of_le := insertIdx_eraseIdx_of_le +@[deprecated (since := "2024-10-21")] alias insertNth_comm := insertIdx_comm +@[deprecated (since := "2024-10-21")] alias mem_insertNth := mem_insertIdx +@[deprecated (since := "2024-10-21")] alias insertNth_of_length_lt := insertIdx_of_length_lt +@[deprecated (since := "2024-10-21")] alias insertNth_length_self := insertIdx_length_self +@[deprecated (since := "2024-10-21")] alias length_le_length_insertNth := length_le_length_insertIdx +@[deprecated (since := "2024-10-21")] alias length_insertNth_le_succ := length_insertIdx_le_succ +@[deprecated (since := "2024-10-21")] alias getElem_insertNth_of_lt := getElem_insertIdx_of_lt +@[deprecated (since := "2024-10-21")] alias get_insertNth_of_lt := get_insertIdx_of_lt +@[deprecated (since := "2024-10-21")] alias getElem_insertNth_self := getElem_insertIdx_self +@[deprecated (since := "2024-10-21")] alias get_insertNth_self := get_insertIdx_self +@[deprecated (since := "2024-10-21")] alias getElem_insertNth_add_succ := getElem_insertIdx_add_succ +@[deprecated (since := "2024-10-21")] alias get_insertNth_add_succ := get_insertIdx_add_succ +@[deprecated (since := "2024-10-21")] alias insertNth_injective := insertIdx_injective + +end InsertIdx + +end List diff --git a/Mathlib/Data/List/InsertNth.lean b/Mathlib/Data/List/InsertNth.lean index 83437b99c7356..48923ca75aef6 100644 --- a/Mathlib/Data/List/InsertNth.lean +++ b/Mathlib/Data/List/InsertNth.lean @@ -1,189 +1,13 @@ /- -Copyright (c) 2014 Parikshit Khanna. All rights reserved. +Copyright (c) 2024 Lean FRO. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Parikshit Khanna, Jeremy Avigad, Leonardo de Moura, Floris van Doorn, Mario Carneiro +Authors: Kim Morrison -/ -import Mathlib.Data.List.Basic +import Mathlib.Data.List.InsertIdx /-! -# insertNth +This is a stub file for importing `Mathlib.Data.List.InsertNth`, +which has been renamed to `Mathlib.Data.List.InsertIdx`. -Proves various lemmas about `List.insertNth`. +This file can be removed once the deprecation for `List.insertNth` is removed. -/ - -assert_not_exists Set.range - -open Function - -open Nat hiding one_pos - -namespace List - -universe u - -variable {α : Type u} - -section InsertNth - -variable {a : α} - -@[simp] -theorem insertNth_zero (s : List α) (x : α) : insertNth 0 x s = x :: s := - rfl - -@[simp] -theorem insertNth_succ_nil (n : ℕ) (a : α) : insertNth (n + 1) a [] = [] := - rfl - -@[simp] -theorem insertNth_succ_cons (s : List α) (hd x : α) (n : ℕ) : - insertNth (n + 1) x (hd :: s) = hd :: insertNth n x s := - rfl - -theorem length_insertNth : ∀ n as, n ≤ length as → length (insertNth n a as) = length as + 1 - | 0, _, _ => rfl - | _ + 1, [], h => (Nat.not_succ_le_zero _ h).elim - | n + 1, _ :: as, h => congr_arg Nat.succ <| length_insertNth n as (Nat.le_of_succ_le_succ h) - -theorem eraseIdx_insertNth (n : ℕ) (l : List α) : (l.insertNth n a).eraseIdx n = l := by - rw [eraseIdx_eq_modifyNthTail, insertNth, modifyNthTail_modifyNthTail_same] - exact modifyNthTail_id _ _ - -@[deprecated (since := "2024-05-04")] alias removeNth_insertNth := eraseIdx_insertNth - -theorem insertNth_eraseIdx_of_ge : - ∀ n m as, - n < length as → n ≤ m → insertNth m a (as.eraseIdx n) = (as.insertNth (m + 1) a).eraseIdx n - | 0, 0, [], has, _ => (lt_irrefl _ has).elim - | 0, 0, _ :: as, _, _ => by simp [eraseIdx, insertNth] - | 0, _ + 1, _ :: _, _, _ => rfl - | n + 1, m + 1, a :: as, has, hmn => - congr_arg (cons a) <| - insertNth_eraseIdx_of_ge n m as (Nat.lt_of_succ_lt_succ has) (Nat.le_of_succ_le_succ hmn) - -@[deprecated (since := "2024-05-04")] alias insertNth_removeNth_of_ge := insertNth_eraseIdx_of_ge - -theorem insertNth_eraseIdx_of_le : - ∀ n m as, - n < length as → m ≤ n → insertNth m a (as.eraseIdx n) = (as.insertNth m a).eraseIdx (n + 1) - | _, 0, _ :: _, _, _ => rfl - | n + 1, m + 1, a :: as, has, hmn => - congr_arg (cons a) <| - insertNth_eraseIdx_of_le n m as (Nat.lt_of_succ_lt_succ has) (Nat.le_of_succ_le_succ hmn) - -@[deprecated (since := "2024-05-04")] alias insertNth_removeNth_of_le := insertNth_eraseIdx_of_le - -theorem insertNth_comm (a b : α) : - ∀ (i j : ℕ) (l : List α) (_ : i ≤ j) (_ : j ≤ length l), - (l.insertNth i a).insertNth (j + 1) b = (l.insertNth j b).insertNth i a - | 0, j, l => by simp [insertNth] - | _ + 1, 0, _ => fun h => (Nat.not_lt_zero _ h).elim - | i + 1, j + 1, [] => by simp - | i + 1, j + 1, c :: l => fun h₀ h₁ => by - simp only [insertNth_succ_cons, cons.injEq, true_and] - exact insertNth_comm a b i j l (Nat.le_of_succ_le_succ h₀) (Nat.le_of_succ_le_succ h₁) - -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefixes below. --/ -theorem mem_insertNth {a b : α} : - ∀ {n : ℕ} {l : List α} (_ : n ≤ l.length), a ∈ l.insertNth n b ↔ a = b ∨ a ∈ l - | 0, as, _ => by simp - | _ + 1, [], h => (Nat.not_succ_le_zero _ h).elim - | n + 1, a' :: as, h => by - rw [List.insertNth_succ_cons, mem_cons, mem_insertNth (Nat.le_of_succ_le_succ h), - ← _root_.or_assoc, @or_comm (a = a'), _root_.or_assoc, mem_cons] - -theorem insertNth_of_length_lt (l : List α) (x : α) (n : ℕ) (h : l.length < n) : - insertNth n x l = l := by - induction' l with hd tl IH generalizing n - · cases n - · simp at h - · simp - · cases n - · simp at h - · simp only [Nat.succ_lt_succ_iff, length] at h - simpa using IH _ h - -@[simp] -theorem insertNth_length_self (l : List α) (x : α) : insertNth l.length x l = l ++ [x] := by - induction' l with hd tl IH - · simp - · simpa using IH - -theorem length_le_length_insertNth (l : List α) (x : α) (n : ℕ) : - l.length ≤ (insertNth n x l).length := by - rcases le_or_lt n l.length with hn | hn - · rw [length_insertNth _ _ hn] - exact (Nat.lt_succ_self _).le - · rw [insertNth_of_length_lt _ _ _ hn] - -theorem length_insertNth_le_succ (l : List α) (x : α) (n : ℕ) : - (insertNth n x l).length ≤ l.length + 1 := by - rcases le_or_lt n l.length with hn | hn - · rw [length_insertNth _ _ hn] - · rw [insertNth_of_length_lt _ _ _ hn] - exact (Nat.lt_succ_self _).le - -theorem getElem_insertNth_of_lt (l : List α) (x : α) (n k : ℕ) (hn : k < n) (hk : k < l.length) - (hk' : k < (insertNth n x l).length := hk.trans_le (length_le_length_insertNth _ _ _)) : - (insertNth n x l)[k] = l[k] := by - induction' n with n IH generalizing k l - · simp at hn - · cases' l with hd tl - · simp - · cases k - · simp [get] - · rw [Nat.succ_lt_succ_iff] at hn - simpa using IH _ _ hn _ - -theorem get_insertNth_of_lt (l : List α) (x : α) (n k : ℕ) (hn : k < n) (hk : k < l.length) - (hk' : k < (insertNth n x l).length := hk.trans_le (length_le_length_insertNth _ _ _)) : - (insertNth n x l).get ⟨k, hk'⟩ = l.get ⟨k, hk⟩ := by - simp_all [getElem_insertNth_of_lt] - -@[simp] -theorem getElem_insertNth_self (l : List α) (x : α) (n : ℕ) (hn : n ≤ l.length) - (hn' : n < (insertNth n x l).length := (by rwa [length_insertNth _ _ hn, Nat.lt_succ_iff])) : - (insertNth n x l)[n] = x := by - induction' l with hd tl IH generalizing n - · simp only [length] at hn - cases hn - simp only [insertNth_zero, getElem_singleton] - · cases n - · simp - · simp only [Nat.succ_le_succ_iff, length] at hn - simpa using IH _ hn - -theorem get_insertNth_self (l : List α) (x : α) (n : ℕ) (hn : n ≤ l.length) - (hn' : n < (insertNth n x l).length := (by rwa [length_insertNth _ _ hn, Nat.lt_succ_iff])) : - (insertNth n x l).get ⟨n, hn'⟩ = x := by - simp [hn, hn'] - -theorem getElem_insertNth_add_succ (l : List α) (x : α) (n k : ℕ) (hk' : n + k < l.length) - (hk : n + k + 1 < (insertNth n x l).length := (by - rwa [length_insertNth _ _ (by omega), Nat.succ_lt_succ_iff])) : - (insertNth n x l)[n + k + 1] = l[n + k] := by - induction' l with hd tl IH generalizing n k - · simp at hk' - · cases n - · simp - · simpa [Nat.add_right_comm] using IH _ _ _ - -theorem get_insertNth_add_succ (l : List α) (x : α) (n k : ℕ) (hk' : n + k < l.length) - (hk : n + k + 1 < (insertNth n x l).length := (by - rwa [length_insertNth _ _ (by omega), Nat.succ_lt_succ_iff])) : - (insertNth n x l).get ⟨n + k + 1, hk⟩ = get l ⟨n + k, hk'⟩ := by - simp [getElem_insertNth_add_succ, hk, hk'] - -set_option linter.unnecessarySimpa false in -theorem insertNth_injective (n : ℕ) (x : α) : Function.Injective (insertNth n x) := by - induction' n with n IH - · have : insertNth 0 x = cons x := funext fun _ => rfl - simp [this] - · rintro (_ | ⟨a, as⟩) (_ | ⟨b, bs⟩) h <;> simpa [IH.eq_iff] using h - -end InsertNth - -end List diff --git a/Mathlib/Data/List/Lemmas.lean b/Mathlib/Data/List/Lemmas.lean index 3c8d444931883..50f55ed3f3832 100644 --- a/Mathlib/Data/List/Lemmas.lean +++ b/Mathlib/Data/List/Lemmas.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yakov Pechersky, Yury Kudryashov -/ import Mathlib.Data.Set.Image -import Mathlib.Data.List.InsertNth +import Mathlib.Data.List.InsertIdx /-! # Some lemmas about lists involving sets @@ -32,8 +32,8 @@ theorem tail_reverse_eq_reverse_dropLast (l : List α) : @[deprecated (since := "2024-08-19")] alias nthLe_tail := getElem_tail -theorem injOn_insertNth_index_of_not_mem (l : List α) (x : α) (hx : x ∉ l) : - Set.InjOn (fun k => insertNth k x l) { n | n ≤ l.length } := by +theorem injOn_insertIdx_index_of_not_mem (l : List α) (x : α) (hx : x ∉ l) : + Set.InjOn (fun k => insertIdx k x l) { n | n ≤ l.length } := by induction' l with hd tl IH · intro n hn m hm _ simp_all [Set.mem_singleton_iff, Set.setOf_eq_eq_singleton, length] @@ -44,12 +44,15 @@ theorem injOn_insertNth_index_of_not_mem (l : List α) (x : α) (hx : x ∉ l) : · rfl · simp [hx.left] at h · simp [Ne.symm hx.left] at h - · simp only [true_and, eq_self_iff_true, insertNth_succ_cons] at h + · simp only [true_and, eq_self_iff_true, insertIdx_succ_cons] at h rw [Nat.succ_inj'] refine IH hx.right ?_ ?_ (by injection h) · simpa [Nat.succ_le_succ_iff] using hn · simpa [Nat.succ_le_succ_iff] using hm +@[deprecated (since := "2024-10-21")] +alias injOn_insertNth_index_of_not_mem := injOn_insertIdx_index_of_not_mem + theorem foldr_range_subset_of_range_subset {f : β → α → α} {g : γ → α → α} (hfg : Set.range f ⊆ Set.range g) (a : α) : Set.range (foldr f a) ⊆ Set.range (foldr g a) := by rintro _ ⟨l, rfl⟩ diff --git a/Mathlib/Data/List/Permutation.lean b/Mathlib/Data/List/Permutation.lean index f03f62f158c49..5bcb74097b405 100644 --- a/Mathlib/Data/List/Permutation.lean +++ b/Mathlib/Data/List/Permutation.lean @@ -8,7 +8,7 @@ import Mathlib.Data.Nat.Factorial.Basic import Mathlib.Data.List.Count import Mathlib.Data.List.Dedup import Mathlib.Data.List.Duplicate -import Mathlib.Data.List.InsertNth +import Mathlib.Data.List.InsertIdx import Mathlib.Data.List.Lattice import Mathlib.Data.List.Perm import Batteries.Data.List.Perm @@ -390,7 +390,7 @@ theorem perm_permutations'_iff {s t : List α} : permutations' s ~ permutations' theorem getElem_permutations'Aux (s : List α) (x : α) (n : ℕ) (hn : n < length (permutations'Aux x s)) : - (permutations'Aux x s)[n] = s.insertNth n x := by + (permutations'Aux x s)[n] = s.insertIdx n x := by induction' s with y s IH generalizing n · simp only [length, Nat.zero_add, Nat.lt_one_iff] at hn simp [hn] @@ -400,7 +400,7 @@ theorem getElem_permutations'Aux (s : List α) (x : α) (n : ℕ) theorem get_permutations'Aux (s : List α) (x : α) (n : ℕ) (hn : n < length (permutations'Aux x s)) : - (permutations'Aux x s).get ⟨n, hn⟩ = s.insertNth n x := by + (permutations'Aux x s).get ⟨n, hn⟩ = s.insertIdx n x := by simp [getElem_permutations'Aux] theorem count_permutations'Aux_self [DecidableEq α] (l : List α) (x : α) : @@ -430,7 +430,7 @@ theorem permutations'Aux_get_zero (s : List α) (x : α) theorem injective_permutations'Aux (x : α) : Function.Injective (permutations'Aux x) := by intro s t h - apply insertNth_injective s.length x + apply insertIdx_injective s.length x have hl : s.length = t.length := by simpa using congr_arg length h rw [← get_permutations'Aux s x s.length (by simp), ← get_permutations'Aux t x s.length (by simp [hl])] @@ -456,21 +456,21 @@ theorem nodup_permutations'Aux_iff {s : List α} {x : α} : Nodup (permutations' have k1l : k + 1 < (permutations'Aux x s).length := by simpa using hk rw [← @Fin.mk.inj_iff _ _ _ kl k1l]; apply h rw [get_permutations'Aux, get_permutations'Aux] - have hl : length (insertNth k x s) = length (insertNth (k + 1) x s) := by - rw [length_insertNth _ _ hk.le, length_insertNth _ _ (Nat.succ_le_of_lt hk)] + have hl : length (insertIdx k x s) = length (insertIdx (k + 1) x s) := by + rw [length_insertIdx _ _ hk.le, length_insertIdx _ _ (Nat.succ_le_of_lt hk)] refine ext_get hl fun n hn hn' => ?_ rcases lt_trichotomy n k with (H | rfl | H) - · rw [get_insertNth_of_lt _ _ _ _ H (H.trans hk), - get_insertNth_of_lt _ _ _ _ (H.trans (Nat.lt_succ_self _))] - · rw [get_insertNth_self _ _ _ hk.le, get_insertNth_of_lt _ _ _ _ (Nat.lt_succ_self _) hk, hk'] + · rw [get_insertIdx_of_lt _ _ _ _ H (H.trans hk), + get_insertIdx_of_lt _ _ _ _ (H.trans (Nat.lt_succ_self _))] + · rw [get_insertIdx_self _ _ _ hk.le, get_insertIdx_of_lt _ _ _ _ (Nat.lt_succ_self _) hk, hk'] · rcases (Nat.succ_le_of_lt H).eq_or_lt with (rfl | H') - · rw [get_insertNth_self _ _ _ (Nat.succ_le_of_lt hk)] + · rw [get_insertIdx_self _ _ _ (Nat.succ_le_of_lt hk)] convert hk' using 1 - exact get_insertNth_add_succ _ _ _ 0 _ + exact get_insertIdx_add_succ _ _ _ 0 _ · obtain ⟨m, rfl⟩ := Nat.exists_eq_add_of_lt H' - rw [length_insertNth _ _ hk.le, Nat.succ_lt_succ_iff, Nat.succ_add] at hn - rw [get_insertNth_add_succ] - · convert get_insertNth_add_succ s x k m.succ (by simpa using hn) using 2 + rw [length_insertIdx _ _ hk.le, Nat.succ_lt_succ_iff, Nat.succ_add] at hn + rw [get_insertIdx_add_succ] + · convert get_insertIdx_add_succ s x k m.succ (by simpa using hn) using 2 · simp [Nat.add_assoc, Nat.add_left_comm] · simp [Nat.add_left_comm, Nat.add_comm] · simpa [Nat.succ_add] using hn @@ -496,20 +496,20 @@ theorem nodup_permutations (s : List α) (hs : Nodup s) : Nodup s.permutations : simp only [Nat.lt_succ_iff, length_permutations'Aux] at hn hm rw [get_permutations'Aux] at hn' hm' have hx : - (insertNth n x as)[m]'(by rwa [length_insertNth _ _ hn, Nat.lt_succ_iff, hl]) = x := by + (insertIdx n x as)[m]'(by rwa [length_insertIdx _ _ hn, Nat.lt_succ_iff, hl]) = x := by simp [hn', ← hm', hm] have hx' : - (insertNth m x bs)[n]'(by rwa [length_insertNth _ _ hm, Nat.lt_succ_iff, ← hl]) = x := by + (insertIdx m x bs)[n]'(by rwa [length_insertIdx _ _ hm, Nat.lt_succ_iff, ← hl]) = x := by simp [hm', ← hn', hn] rcases lt_trichotomy n m with (ht | ht | ht) · suffices x ∈ bs by exact h x (hb.subset this) rfl - rw [← hx', getElem_insertNth_of_lt _ _ _ _ ht (ht.trans_le hm)] + rw [← hx', getElem_insertIdx_of_lt _ _ _ _ ht (ht.trans_le hm)] exact get_mem _ _ _ · simp only [ht] at hm' hn' rw [← hm'] at hn' - exact H (insertNth_injective _ _ hn') + exact H (insertIdx_injective _ _ hn') · suffices x ∈ as by exact h x (ha.subset this) rfl - rw [← hx, getElem_insertNth_of_lt _ _ _ _ ht (ht.trans_le hn)] + rw [← hx, getElem_insertIdx_of_lt _ _ _ _ ht (ht.trans_le hn)] exact get_mem _ _ _ lemma permutations_take_two (x y : α) (s : List α) : diff --git a/Mathlib/Data/Vector/Basic.lean b/Mathlib/Data/Vector/Basic.lean index 67059f9ba1cbd..5e7c6ea9c6715 100644 --- a/Mathlib/Data/Vector/Basic.lean +++ b/Mathlib/Data/Vector/Basic.lean @@ -7,7 +7,7 @@ import Mathlib.Algebra.BigOperators.Group.List import Mathlib.Data.Vector.Defs import Mathlib.Data.List.Nodup import Mathlib.Data.List.OfFn -import Mathlib.Data.List.InsertNth +import Mathlib.Data.List.InsertIdx import Mathlib.Control.Applicative import Mathlib.Control.Traversable.Basic @@ -472,71 +472,81 @@ def casesOn₃ {motive : ∀{n}, Vector α n → Vector β n → Vector γ n def toArray : Vector α n → Array α | ⟨xs, _⟩ => cast (by rfl) xs.toArray -section InsertNth +section InsertIdx variable {a : α} -/-- `v.insertNth a i` inserts `a` into the vector `v` at position `i` +/-- `v.insertIdx a i` inserts `a` into the vector `v` at position `i` (and shifting later components to the right). -/ -def insertNth (a : α) (i : Fin (n + 1)) (v : Vector α n) : Vector α (n + 1) := - ⟨v.1.insertNth i a, by - rw [List.length_insertNth, v.2] +def insertIdx (a : α) (i : Fin (n + 1)) (v : Vector α n) : Vector α (n + 1) := + ⟨v.1.insertIdx i a, by + rw [List.length_insertIdx, v.2] rw [v.2, ← Nat.succ_le_succ_iff] exact i.2⟩ -theorem insertNth_val {i : Fin (n + 1)} {v : Vector α n} : - (v.insertNth a i).val = v.val.insertNth i.1 a := +@[deprecated (since := "2024-10-21")] alias insertNth := insertIdx + +theorem insertIdx_val {i : Fin (n + 1)} {v : Vector α n} : + (v.insertIdx a i).val = v.val.insertIdx i.1 a := rfl +@[deprecated (since := "2024-10-21")] alias insertNth_val := insertIdx_val + @[simp] theorem eraseIdx_val {i : Fin n} : ∀ {v : Vector α n}, (eraseIdx i v).val = v.val.eraseIdx i | _ => rfl +@[deprecated (since := "2024-10-21")] alias eraseNth_val := eraseIdx_val @[deprecated (since := "2024-05-04")] alias removeNth_val := eraseIdx_val -theorem eraseIdx_insertNth {v : Vector α n} {i : Fin (n + 1)} : - eraseIdx i (insertNth a i v) = v := - Subtype.eq <| List.eraseIdx_insertNth i.1 v.1 +theorem eraseIdx_insertIdx {v : Vector α n} {i : Fin (n + 1)} : + eraseIdx i (insertIdx a i v) = v := + Subtype.eq <| List.eraseIdx_insertIdx i.1 v.1 -@[deprecated (since := "2024-05-04")] alias removeNth_insertNth := eraseIdx_insertNth +@[deprecated (since := "2024-05-04")] alias eraseIdx_insertNth := eraseIdx_insertIdx +@[deprecated (since := "2024-05-04")] alias removeNth_insertNth := eraseIdx_insertIdx -theorem eraseIdx_insertNth' {v : Vector α (n + 1)} : +/-- Erasing an element after inserting an element, at different indices. -/ +theorem eraseIdx_insertIdx' {v : Vector α (n + 1)} : ∀ {i : Fin (n + 1)} {j : Fin (n + 2)}, - eraseIdx (j.succAbove i) (insertNth a j v) = insertNth a (i.predAbove j) (eraseIdx i v) + eraseIdx (j.succAbove i) (insertIdx a j v) = insertIdx a (i.predAbove j) (eraseIdx i v) | ⟨i, hi⟩, ⟨j, hj⟩ => by - dsimp [insertNth, eraseIdx, Fin.succAbove, Fin.predAbove] + dsimp [insertIdx, eraseIdx, Fin.succAbove, Fin.predAbove] rw [Subtype.mk_eq_mk] simp only [Fin.lt_iff_val_lt_val] split_ifs with hij · rcases Nat.exists_eq_succ_of_ne_zero (Nat.pos_iff_ne_zero.1 (lt_of_le_of_lt (Nat.zero_le _) hij)) with ⟨j, rfl⟩ - rw [← List.insertNth_eraseIdx_of_ge] + rw [← List.insertIdx_eraseIdx_of_ge] · simp; rfl · simpa · simpa [Nat.lt_succ_iff] using hij · dsimp - rw [← List.insertNth_eraseIdx_of_le i j _ _ _] + rw [← List.insertIdx_eraseIdx_of_le i j _ _ _] · rfl · simpa · simpa [not_lt] using hij -@[deprecated (since := "2024-05-04")] alias removeNth_insertNth' := eraseIdx_insertNth' +@[deprecated (since := "2024-05-04")] alias eraseIdx_insertNth' := eraseIdx_insertIdx' +@[deprecated (since := "2024-05-04")] alias removeNth_insertNth' := eraseIdx_insertIdx' -theorem insertNth_comm (a b : α) (i j : Fin (n + 1)) (h : i ≤ j) : +theorem insertIdx_comm (a b : α) (i j : Fin (n + 1)) (h : i ≤ j) : ∀ v : Vector α n, - (v.insertNth a i).insertNth b j.succ = (v.insertNth b j).insertNth a (Fin.castSucc i) + (v.insertIdx a i).insertIdx b j.succ = (v.insertIdx b j).insertIdx a (Fin.castSucc i) | ⟨l, hl⟩ => by refine Subtype.eq ?_ - simp only [insertNth_val, Fin.val_succ, Fin.castSucc, Fin.coe_castAdd] - apply List.insertNth_comm + simp only [insertIdx_val, Fin.val_succ, Fin.castSucc, Fin.coe_castAdd] + apply List.insertIdx_comm · assumption · rw [hl] exact Nat.le_of_succ_le_succ j.2 -end InsertNth +@[deprecated (since := "2024-10-21")] alias insertNth_comm := insertIdx_comm + +end InsertIdx -- Porting note: renamed to `set` from `updateNth` to align with `List` -section ModifyNth +section Set /-- `set v n a` replaces the `n`th element of `v` with `a`. -/ def set (v : Vector α n) (i : Fin n) (a : α) : Vector α n := @@ -574,7 +584,7 @@ theorem prod_set' [CommGroup α] (v : Vector α n) (i : Fin n) (a : α) : refine (List.prod_set' v.toList i a).trans ?_ simp [get_eq_get, mul_assoc] -end ModifyNth +end Set end Vector diff --git a/Mathlib/SetTheory/Game/PGame.lean b/Mathlib/SetTheory/Game/PGame.lean index 83f9d1f8d2559..56ee1377443a3 100644 --- a/Mathlib/SetTheory/Game/PGame.lean +++ b/Mathlib/SetTheory/Game/PGame.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Reid Barton, Mario Carneiro, Isabel Longbottom, Kim Morrison -/ import Mathlib.Algebra.Order.ZeroLEOne -import Mathlib.Data.List.InsertNth +import Mathlib.Data.List.InsertIdx import Mathlib.Logic.Relation import Mathlib.Logic.Small.Defs import Mathlib.Order.GameAdd diff --git a/Mathlib/Testing/SlimCheck/Gen.lean b/Mathlib/Testing/SlimCheck/Gen.lean index 0f3d66842989e..8429d8cbfff50 100644 --- a/Mathlib/Testing/SlimCheck/Gen.lean +++ b/Mathlib/Testing/SlimCheck/Gen.lean @@ -101,7 +101,7 @@ def permutationOf : (xs : List α) → Gen { ys // xs ~ ys } | x::xs => do let ⟨ys, h1⟩ ← permutationOf xs let ⟨n, _, h3⟩ ← ULiftable.up <| choose Nat 0 ys.length (Nat.zero_le _) - pure ⟨insertNth n x ys, Perm.trans (Perm.cons _ h1) (perm_insertNth _ _ h3).symm⟩ + pure ⟨insertIdx n x ys, Perm.trans (Perm.cons _ h1) (perm_insertIdx _ _ h3).symm⟩ /-- Given two generators produces a tuple consisting out of the result of both -/ def prodOf {α : Type u} {β : Type v} (x : Gen α) (y : Gen β) : Gen (α × β) := do diff --git a/Mathlib/Testing/SlimCheck/Sampleable.lean b/Mathlib/Testing/SlimCheck/Sampleable.lean index 96f6e2b638df6..cb9e8a50dcfb2 100644 --- a/Mathlib/Testing/SlimCheck/Sampleable.lean +++ b/Mathlib/Testing/SlimCheck/Sampleable.lean @@ -188,7 +188,7 @@ open Shrinkable instance List.shrinkable [Shrinkable α] : Shrinkable (List α) where shrink := fun L => (L.mapIdx fun i _ => L.eraseIdx i) ++ - (L.mapIdx fun i a => (shrink a).map fun a' => L.modifyNth (fun _ => a') i).join + (L.mapIdx fun i a => (shrink a).map fun a' => L.modify (fun _ => a') i).join end Shrinkers diff --git a/Mathlib/Topology/List.lean b/Mathlib/Topology/List.lean index ca954e2c62243..fd9df79a1b592 100644 --- a/Mathlib/Topology/List.lean +++ b/Mathlib/Topology/List.lean @@ -115,9 +115,10 @@ theorem continuousAt_length : ∀ l : List α, ContinuousAt List.length l := by refine Tendsto.comp (tendsto_pure_pure (fun x => x + 1) _) ?_ exact Tendsto.comp ih tendsto_snd -theorem tendsto_insertNth' {a : α} : +/-- Continuity of `insertIdx` in terms of `Tendsto`. -/ +theorem tendsto_insertIdx' {a : α} : ∀ {n : ℕ} {l : List α}, - Tendsto (fun p : α × List α => insertNth n p.1 p.2) (𝓝 a ×ˢ 𝓝 l) (𝓝 (insertNth n a l)) + Tendsto (fun p : α × List α => insertIdx n p.1 p.2) (𝓝 a ×ˢ 𝓝 l) (𝓝 (insertIdx n a l)) | 0, _ => tendsto_cons | n + 1, [] => by simp | n + 1, a'::l => by @@ -128,16 +129,22 @@ theorem tendsto_insertNth' {a : α} : rw [this, tendsto_map'_iff] exact (tendsto_fst.comp tendsto_snd).cons - ((@tendsto_insertNth' _ n l).comp <| tendsto_fst.prod_mk <| tendsto_snd.comp tendsto_snd) + ((@tendsto_insertIdx' _ n l).comp <| tendsto_fst.prod_mk <| tendsto_snd.comp tendsto_snd) -theorem tendsto_insertNth {β} {n : ℕ} {a : α} {l : List α} {f : β → α} {g : β → List α} +@[deprecated (since := "2024-10-21")] alias tendsto_insertNth' := tendsto_insertIdx' + +theorem tendsto_insertIdx {β} {n : ℕ} {a : α} {l : List α} {f : β → α} {g : β → List α} {b : Filter β} (hf : Tendsto f b (𝓝 a)) (hg : Tendsto g b (𝓝 l)) : - Tendsto (fun b : β => insertNth n (f b) (g b)) b (𝓝 (insertNth n a l)) := - tendsto_insertNth'.comp (Tendsto.prod_mk hf hg) + Tendsto (fun b : β => insertIdx n (f b) (g b)) b (𝓝 (insertIdx n a l)) := + tendsto_insertIdx'.comp (Tendsto.prod_mk hf hg) + +@[deprecated (since := "2024-10-21")] alias tendsto_insertNth := tendsto_insertIdx' -theorem continuous_insertNth {n : ℕ} : Continuous fun p : α × List α => insertNth n p.1 p.2 := +theorem continuous_insertIdx {n : ℕ} : Continuous fun p : α × List α => insertIdx n p.1 p.2 := continuous_iff_continuousAt.mpr fun ⟨a, l⟩ => by - rw [ContinuousAt, nhds_prod_eq]; exact tendsto_insertNth' + rw [ContinuousAt, nhds_prod_eq]; exact tendsto_insertIdx' + +@[deprecated (since := "2024-10-21")] alias continuous_insertNth := continuous_insertIdx theorem tendsto_eraseIdx : ∀ {n : ℕ} {l : List α}, Tendsto (eraseIdx · n) (𝓝 l) (𝓝 (eraseIdx l n)) @@ -184,23 +191,30 @@ theorem tendsto_cons {n : ℕ} {a : α} {l : Vector α n} : rw [tendsto_subtype_rng, Vector.cons_val] exact tendsto_fst.cons (Tendsto.comp continuousAt_subtype_val tendsto_snd) -theorem tendsto_insertNth {n : ℕ} {i : Fin (n + 1)} {a : α} : +theorem tendsto_insertIdx {n : ℕ} {i : Fin (n + 1)} {a : α} : ∀ {l : Vector α n}, - Tendsto (fun p : α × Vector α n => Vector.insertNth p.1 i p.2) (𝓝 a ×ˢ 𝓝 l) - (𝓝 (Vector.insertNth a i l)) + Tendsto (fun p : α × Vector α n => Vector.insertIdx p.1 i p.2) (𝓝 a ×ˢ 𝓝 l) + (𝓝 (Vector.insertIdx a i l)) | ⟨l, hl⟩ => by - rw [Vector.insertNth, tendsto_subtype_rng] - simp only [Vector.insertNth_val] - exact List.tendsto_insertNth tendsto_fst (Tendsto.comp continuousAt_subtype_val tendsto_snd : _) + rw [Vector.insertIdx, tendsto_subtype_rng] + simp only [Vector.insertIdx_val] + exact List.tendsto_insertIdx tendsto_fst (Tendsto.comp continuousAt_subtype_val tendsto_snd : _) -theorem continuous_insertNth' {n : ℕ} {i : Fin (n + 1)} : - Continuous fun p : α × Vector α n => Vector.insertNth p.1 i p.2 := +@[deprecated (since := "2024-10-21")] alias tendsto_insertNth := tendsto_insertIdx' + +/-- Continuity of `Vector.insertIdx`. -/ +theorem continuous_insertIdx' {n : ℕ} {i : Fin (n + 1)} : + Continuous fun p : α × Vector α n => Vector.insertIdx p.1 i p.2 := continuous_iff_continuousAt.mpr fun ⟨a, l⟩ => by - rw [ContinuousAt, nhds_prod_eq]; exact tendsto_insertNth + rw [ContinuousAt, nhds_prod_eq]; exact tendsto_insertIdx + +@[deprecated (since := "2024-10-21")] alias continuous_insertNth' := continuous_insertIdx' + +theorem continuous_insertIdx {n : ℕ} {i : Fin (n + 1)} {f : β → α} {g : β → Vector α n} + (hf : Continuous f) (hg : Continuous g) : Continuous fun b => Vector.insertIdx (f b) i (g b) := + continuous_insertIdx'.comp (hf.prod_mk hg : _) -theorem continuous_insertNth {n : ℕ} {i : Fin (n + 1)} {f : β → α} {g : β → Vector α n} - (hf : Continuous f) (hg : Continuous g) : Continuous fun b => Vector.insertNth (f b) i (g b) := - continuous_insertNth'.comp (hf.prod_mk hg : _) +@[deprecated (since := "2024-10-21")] alias continuous_insertNth := continuous_insertIdx theorem continuousAt_eraseIdx {n : ℕ} {i : Fin (n + 1)} : ∀ {l : Vector α (n + 1)}, ContinuousAt (Vector.eraseIdx i) l diff --git a/lake-manifest.json b/lake-manifest.json index 8d8fe8356466a..21a3f906a6a89 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,10 +5,10 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "10130798199d306703dee5ab2567961444ebbd04", + "rev": "50b1bcef8cde122ad51182ec78ec8cb5ba37c60f", "name": "batteries", "manifestFile": "lake-manifest.json", - "inputRev": "main", + "inputRev": "mv_modifyNth", "inherited": false, "configFile": "lakefile.lean"}, {"url": "https://github.com/leanprover-community/quote4", diff --git a/lakefile.lean b/lakefile.lean index 50c310fac84a6..3d8878457a94f 100644 --- a/lakefile.lean +++ b/lakefile.lean @@ -7,7 +7,7 @@ open Lake DSL ## Mathlib dependencies on upstream projects -/ -require "leanprover-community" / "batteries" @ git "main" +require "leanprover-community" / "batteries" @ git "mv_modifyNth" require "leanprover-community" / "Qq" @ git "master" require "leanprover-community" / "aesop" @ git "master" require "leanprover-community" / "proofwidgets" @ git "v0.0.43" diff --git a/scripts/noshake.json b/scripts/noshake.json index 5ee9b18083934..22149664826c0 100644 --- a/scripts/noshake.json +++ b/scripts/noshake.json @@ -364,7 +364,7 @@ "Mathlib.Data.List.Range": ["Batteries.Data.Nat.Lemmas"], "Mathlib.Data.List.ProdSigma": ["Batteries.Data.Nat.Lemmas", "Mathlib.Data.List.Basic"], - "Mathlib.Data.List.Lemmas": ["Mathlib.Data.List.InsertNth"], + "Mathlib.Data.List.Lemmas": ["Mathlib.Data.List.InsertIdx"], "Mathlib.Data.List.Defs": ["Batteries.Data.RBMap.Basic"], "Mathlib.Data.List.Basic": ["Mathlib.Control.Basic", "Mathlib.Data.Option.Basic"], diff --git a/test/rewrites.lean b/test/rewrites.lean index bbd4e2eef294f..2b4fe6758c081 100644 --- a/test/rewrites.lean +++ b/test/rewrites.lean @@ -1,6 +1,6 @@ import Mathlib.Data.Nat.Prime.Defs import Mathlib.CategoryTheory.Category.Basic -import Mathlib.Data.List.InsertNth +import Mathlib.Data.List.InsertIdx import Mathlib.Algebra.Group.Basic -- This is partially duplicative with the tests for `rw?` in Lean. From c8f7de855ad546b004a0ec1172f9c294105be0a4 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 21 Oct 2024 05:53:15 +0000 Subject: [PATCH 217/505] chore: move batteries back to main; bump (#17991) Batteries was accidentally moved to another branch when #17985 was merged prematurely. --- lake-manifest.json | 4 ++-- lakefile.lean | 2 +- 2 files changed, 3 insertions(+), 3 deletions(-) diff --git a/lake-manifest.json b/lake-manifest.json index 21a3f906a6a89..2736602d58cb5 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,10 +5,10 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "50b1bcef8cde122ad51182ec78ec8cb5ba37c60f", + "rev": "dd6b1019b5cef990161bf3edfebeb6b0be78044a", "name": "batteries", "manifestFile": "lake-manifest.json", - "inputRev": "mv_modifyNth", + "inputRev": "main", "inherited": false, "configFile": "lakefile.lean"}, {"url": "https://github.com/leanprover-community/quote4", diff --git a/lakefile.lean b/lakefile.lean index 3d8878457a94f..50c310fac84a6 100644 --- a/lakefile.lean +++ b/lakefile.lean @@ -7,7 +7,7 @@ open Lake DSL ## Mathlib dependencies on upstream projects -/ -require "leanprover-community" / "batteries" @ git "mv_modifyNth" +require "leanprover-community" / "batteries" @ git "main" require "leanprover-community" / "Qq" @ git "master" require "leanprover-community" / "aesop" @ git "master" require "leanprover-community" / "proofwidgets" @ git "v0.0.43" From f738d2223ef0c3e25d6a5482c94afb6cf6191ffa Mon Sep 17 00:00:00 2001 From: damiano Date: Mon, 21 Oct 2024 06:32:44 +0000 Subject: [PATCH 218/505] feat: the multiGoal linter (#12339) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit A linter that makes sure that, when multiple goals are present, they are enclosed in `·`s. Adaptations (the order is chronological, there should be no dependency among them): * #12338, * #12361, * #12372, * #12560, * #12834, * #12381, * #12908, * #14939, plus many many more that this comment is too small to contain. Co-authored-by: Michael Rothgang --- Mathlib.lean | 1 + Mathlib/Analysis/Analytic/Within.lean | 1 + Mathlib/Init.lean | 1 + Mathlib/LinearAlgebra/Eigenspace/Basic.lean | 2 +- Mathlib/Tactic.lean | 1 + Mathlib/Tactic/Linter/Multigoal.lean | 165 ++++++++++++++++++++ lakefile.lean | 1 + test/Multigoal.lean | 123 +++++++++++++++ 8 files changed, 294 insertions(+), 1 deletion(-) create mode 100644 Mathlib/Tactic/Linter/Multigoal.lean create mode 100644 test/Multigoal.lean diff --git a/Mathlib.lean b/Mathlib.lean index e722c84da9b0e..7c599ec0aa9d0 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4398,6 +4398,7 @@ import Mathlib.Tactic.Linter.HaveLetLinter import Mathlib.Tactic.Linter.Header import Mathlib.Tactic.Linter.Lint import Mathlib.Tactic.Linter.MinImports +import Mathlib.Tactic.Linter.Multigoal import Mathlib.Tactic.Linter.OldObtain import Mathlib.Tactic.Linter.PPRoundtrip import Mathlib.Tactic.Linter.RefineLinter diff --git a/Mathlib/Analysis/Analytic/Within.lean b/Mathlib/Analysis/Analytic/Within.lean index 7703eb9524ef0..b9f6a8bda1b70 100644 --- a/Mathlib/Analysis/Analytic/Within.lean +++ b/Mathlib/Analysis/Analytic/Within.lean @@ -113,6 +113,7 @@ result for `AnalyticOn`, as this requires a bit more work to show that local ext be stitched together. -/ +set_option linter.style.multiGoal false in /-- `f` has power series `p` at `x` iff some local extension of `f` has that series -/ lemma hasFPowerSeriesWithinOnBall_iff_exists_hasFPowerSeriesOnBall [CompleteSpace F] {f : E → F} {p : FormalMultilinearSeries 𝕜 E F} {s : Set E} {x : E} {r : ℝ≥0∞} : diff --git a/Mathlib/Init.lean b/Mathlib/Init.lean index bffb68e9d4831..32cb637b4877d 100644 --- a/Mathlib/Init.lean +++ b/Mathlib/Init.lean @@ -4,6 +4,7 @@ import Mathlib.Tactic.Linter.GlobalAttributeIn import Mathlib.Tactic.Linter.Header -- This file imports Batteries.Tactic.Lint, where the `env_linter` attribute is defined. import Mathlib.Tactic.Linter.Lint +import Mathlib.Tactic.Linter.Multigoal import Mathlib.Tactic.Linter.OldObtain import Mathlib.Tactic.Linter.RefineLinter import Mathlib.Tactic.Linter.UnusedTactic diff --git a/Mathlib/LinearAlgebra/Eigenspace/Basic.lean b/Mathlib/LinearAlgebra/Eigenspace/Basic.lean index a66604891270f..d323e2e549474 100644 --- a/Mathlib/LinearAlgebra/Eigenspace/Basic.lean +++ b/Mathlib/LinearAlgebra/Eigenspace/Basic.lean @@ -376,7 +376,7 @@ lemma isNilpotent_restrict_unifEigenspace_top [IsNoetherian R M] (f : End R M) ( 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 μ) + on_goal 2 => 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` diff --git a/Mathlib/Tactic.lean b/Mathlib/Tactic.lean index efac05b1b7da1..fe71b6151623d 100644 --- a/Mathlib/Tactic.lean +++ b/Mathlib/Tactic.lean @@ -144,6 +144,7 @@ import Mathlib.Tactic.Linter.HaveLetLinter import Mathlib.Tactic.Linter.Header import Mathlib.Tactic.Linter.Lint import Mathlib.Tactic.Linter.MinImports +import Mathlib.Tactic.Linter.Multigoal import Mathlib.Tactic.Linter.OldObtain import Mathlib.Tactic.Linter.PPRoundtrip import Mathlib.Tactic.Linter.RefineLinter diff --git a/Mathlib/Tactic/Linter/Multigoal.lean b/Mathlib/Tactic/Linter/Multigoal.lean new file mode 100644 index 0000000000000..6d459f2968179 --- /dev/null +++ b/Mathlib/Tactic/Linter/Multigoal.lean @@ -0,0 +1,165 @@ +/- +Copyright (c) 2024 Damiano Testa. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Damiano Testa +-/ +import Lean.Elab.Command + +/-! +# The "multiGoal" linter + +The "multiGoal" linter emits a warning where there is more than a single goal in scope. +There is an exception: a tactic that closes *all* remaining goals is allowed. + +There are a few tactics, such as `skip`, `swap` or the `try` combinator, that are intended to work +specifically in such a situation. +Otherwise, the mathlib style guide ask that goals be be focused until there is only one active goal +at a time. +If this focusing does not happen, the linter emits a warning. +Typically, the focusing is achieved by the `cdot`: `·`, but, e.g., `focus` or `on_goal x` +also serve a similar purpose. + +TODO: +* Should the linter flag unnecessary scoping as well? + For instance, should + ```lean + example : True := by + · · exact .intro + ``` + raise a warning? +* Custom support for "accumulating side-goals", so that once they are all in scope + they can be solved in bulk via `all_goals` or a similar tactic. +* In development, `on_goal` has been partly used as a way of silencing the linter + precisely to allow the accumulation referred to in the previous item. + Maybe revisit usages of `on_goal` and also nested `induction` and `cases`. +-/ + +open Lean Elab + +namespace Mathlib.Linter + +/-- The "multiGoal" linter emits a warning when there are multiple active goals. -/ +register_option linter.style.multiGoal : Bool := { + defValue := false + descr := "enable the multiGoal linter" +} + +namespace Style.multiGoal + +/-- The `SyntaxNodeKind`s in `exclusions` correspond to tactics that the linter allows, +even though there are multiple active goals. +Reasons for admitting a kind in `exclusions` include +* the tactic focuses on one goal, e.g. `·`, `focus`, `on_goal i =>`, ...; +* the tactic is reordering the goals, e.g. `swap`, `rotate_left`, ...; +* the tactic is structuring a proof, e.g. `skip`, `<;>`, ...; +* the tactic is creating new goals, e.g. `constructor`, `cases`, `induction`, .... + +There is some overlap in scope between `ignoreBranch` and `exclusions`. + +Tactic combinators like `repeat` or `try` are a mix of both. +-/ +abbrev exclusions : Std.HashSet SyntaxNodeKind := .ofList [ + -- structuring a proof + ``Lean.Parser.Term.cdot, + ``cdot, + ``cdotTk, + ``Lean.Parser.Tactic.tacticSeqBracketed, + `«;», + `«<;>», + ``Lean.Parser.Tactic.«tactic_<;>_», + `«{», + `«]», + `null, + `then, + `else, + ``Lean.Parser.Tactic.«tacticNext_=>_», + ``Lean.Parser.Tactic.tacticSeq1Indented, + ``Lean.Parser.Tactic.tacticSeq, + -- re-ordering goals + `Batteries.Tactic.tacticSwap, + ``Lean.Parser.Tactic.rotateLeft, + ``Lean.Parser.Tactic.rotateRight, + ``Lean.Parser.Tactic.skip, + `Batteries.Tactic.«tacticOn_goal-_=>_», + `Mathlib.Tactic.«tacticSwap_var__,,», + -- tactic combinators + ``Lean.Parser.Tactic.tacticRepeat_, + ``Lean.Parser.Tactic.tacticTry_, + -- creating new goals + ``Lean.Parser.Tactic.paren, + ``Lean.Parser.Tactic.case, + ``Lean.Parser.Tactic.constructor, + `Mathlib.Tactic.tacticAssumption', + ``Lean.Parser.Tactic.induction, + ``Lean.Parser.Tactic.cases, + ``Lean.Parser.Tactic.intros, + ``Lean.Parser.Tactic.injections, + ``Lean.Parser.Tactic.substVars, + `Batteries.Tactic.«tacticPick_goal-_», + ``Lean.Parser.Tactic.case', + `«tactic#adaptation_note_», + `tacticSleep_heartbeats_ + ] + +/-- The `SyntaxNodeKind`s in `ignoreBranch` correspond to tactics that disable the linter from +their first application until the corresponding proof branch is closed. +Reasons for ignoring these tactics include +* the linter gets confused by the proof management, e.g. `conv`; +* the tactics are *intended* to act on multiple goals, e.g. `repeat`, `any_goals`, `all_goals`, ... + +There is some overlap in scope between `exclusions` and `ignoreBranch`. +-/ +abbrev ignoreBranch : Std.HashSet SyntaxNodeKind := .ofList [ + ``Lean.Parser.Tactic.Conv.conv, + `Mathlib.Tactic.Conv.convLHS, + `Mathlib.Tactic.Conv.convRHS, + ``Lean.Parser.Tactic.first, + ``Lean.Parser.Tactic.repeat', + ``Lean.Parser.Tactic.tacticIterate____, + ``Lean.Parser.Tactic.anyGoals, + ``Lean.Parser.Tactic.allGoals, + ``Lean.Parser.Tactic.focus + ] + +/-- `getManyGoals t` returns the syntax nodes of the `InfoTree` `t` corresponding to tactic calls +which +* leave at least one goal that was not present before it ran; +* are not excluded through `exclusions` or `ignoreBranch`; + +together with the total number of goals +-/ +partial +def getManyGoals : InfoTree → Array (Syntax × Nat) + | .node info args => + let kargs := (args.map getManyGoals).toArray.flatten + if let .ofTacticInfo info := info then + if ignoreBranch.contains info.stx.getKind then #[] else + if let .original .. := info.stx.getHeadInfo then + let newGoals := info.goalsAfter.filter (info.goalsBefore.contains ·) + if newGoals.length != 0 && !exclusions.contains info.stx.getKind then + kargs.push (info.stx, newGoals.length) + else kargs + else kargs + else kargs + | .context _ t => getManyGoals t + | _ => default + +@[inherit_doc Mathlib.Linter.linter.style.multiGoal] +def multiGoalLinter : Linter where run := withSetOptionIn fun _stx ↦ do + unless Linter.getLinterValue linter.style.multiGoal (← getOptions) do + return + if (← get).messages.hasErrors then + return + let trees ← getInfoTrees + for t in trees.toArray do + for (s, n) in getManyGoals t do + Linter.logLint linter.style.multiGoal s + m!"There are {n+1} unclosed goals before '{s}' and \ + at least one remaining goal afterwards.\n\ + Please focus on the current goal, for instance using `·` (typed as \"\\.\")." + +initialize addLinter multiGoalLinter + +end Style.multiGoal + +end Mathlib.Linter diff --git a/lakefile.lean b/lakefile.lean index 50c310fac84a6..697f41bab4e5b 100644 --- a/lakefile.lean +++ b/lakefile.lean @@ -36,6 +36,7 @@ abbrev mathlibOnlyLinters : Array LeanOption := #[ ⟨`linter.style.longLine, true⟩, ⟨`linter.style.longFile, .ofNat 1500⟩, ⟨`linter.style.missingEnd, true⟩, + ⟨`linter.style.multiGoal, true⟩, ⟨`linter.style.setOption, true⟩ ] diff --git a/test/Multigoal.lean b/test/Multigoal.lean new file mode 100644 index 0000000000000..068e62456600e --- /dev/null +++ b/test/Multigoal.lean @@ -0,0 +1,123 @@ +import Mathlib.Tactic.Basic +import Mathlib.Tactic.Conv +import Mathlib.Tactic.Linter.Multigoal +import Mathlib.Util.SleepHeartbeats + +-- The warning generated by `linter.style.multiGoal` is not suppressed by `#guard_msgs`, +-- because the linter is run on `#guard_msgs` itself. This is a known issue, see e.g. +-- https://leanprover.zulipchat.com/#narrow/stream/348111-batteries/topic/unreachableTactic.20linter.20not.20suppressed.20by.20.60.23guard_msgs.60 +-- We jump through an extra hoop here to silence the warning. +set_option linter.style.multiGoal false + +-- A deactivated linter does nothing. +example : True := by + by_cases 0 = 0 + exact .intro + exact .intro + +#guard_msgs(drop warning) in +set_option linter.style.multiGoal true in +/-- +warning: There are 2 unclosed goals before 'exact .intro' and at least one remaining goal afterwards. +Please focus on the current goal, for instance using `·` (typed as "\."). +note: this linter can be disabled with `set_option linter.style.multiGoal false` +-/ +#guard_msgs in +example : True := by + by_cases 0 = 0 + exact .intro + exact .intro + +#guard_msgs(drop warning) in +set_option linter.style.multiGoal true in +/-- +warning: There are 2 unclosed goals before 'assumption' and at least one remaining goal afterwards. +Please focus on the current goal, for instance using `·` (typed as "\."). +note: this linter can be disabled with `set_option linter.style.multiGoal false` +-/ +#guard_msgs in +example {n : Nat} (hn : n = 0) : n + 0 = 0 := by + conv => + congr + rw [← Nat.add_zero 0] + conv_lhs => + congr + rw [← Nat.add_zero n] + rfl + conv_rhs => + rw [← Nat.add_zero 0] + congr + rfl + rfl + by_cases 0 = 0 + assumption + assumption + +set_option linter.unusedTactic false in +#guard_msgs(drop warning) in +set_option linter.style.multiGoal true in +/-- +warning: There are 2 unclosed goals before 'rfl' and at least one remaining goal afterwards. +Please focus on the current goal, for instance using `·` (typed as "\."). +note: this linter can be disabled with `set_option linter.style.multiGoal false` +-/ +#guard_msgs in +example (p : Prop) (hp : p) : (0 = 0 ∧ p) ∨ 0 = 0 := by + iterate left; decide + repeat' left; decide + refine Or.inl ⟨?_, ?_⟩ + rfl + assumption + +#guard_msgs(drop warning) in +set_option linter.style.multiGoal true in +/-- +warning: There are 3 unclosed goals before 'rfl' and at least one remaining goal afterwards. +Please focus on the current goal, for instance using `·` (typed as "\."). +note: this linter can be disabled with `set_option linter.style.multiGoal false` +--- +warning: There are 2 unclosed goals before 'trivial' and at least one remaining goal afterwards. +Please focus on the current goal, for instance using `·` (typed as "\."). +note: this linter can be disabled with `set_option linter.style.multiGoal false` +-/ +#guard_msgs in +example : 0 = 0 ∧ 0 = 0 ∧ 0 = 0 := by + refine ⟨?_, ?_, ?_⟩ + rfl + trivial + rfl + +example (p : Bool) : 0 = 0 := by + cases p + case' false => rfl + case' true => rfl + +#guard_msgs in +-- `assumption'` is allowed, as it is useful precisely when there are multiple active goals +example (p : Bool) (f : False) {h : 0 = 0} : 0 = 0 ∧ 0 = 1 := by + cases p <;> + constructor + assumption' + any_goals cases f + +#guard_msgs in +-- `focus` is ignored +example : True ∧ True := by + constructor + focus + exact .intro + focus + exact .intro + +set_option linter.unusedTactic false in +example : 1 = 1 := by + sleep_heartbeats 1000 + rfl + +-- we test that a tactic closing all remaining goals does not trigger the linter +macro "bi_trivial" : tactic => `(tactic| (trivial; trivial)) + +set_option linter.style.multiGoal true in +example : True ∧ True := by + constructor + bi_trivial From a60b09cd54613409d215c944aebb2cf75991287c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Mon, 21 Oct 2024 06:32:46 +0000 Subject: [PATCH 219/505] =?UTF-8?q?refactor(SetTheory/Ordinal/Arithmetic):?= =?UTF-8?q?=20deprecate=20`blsub=E2=82=82`=20(#17679)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit See the new docstring of `blsub₂` for the reasoning. Co-authored-by: YnirPaz --- Mathlib/SetTheory/Ordinal/Arithmetic.lean | 9 +++--- Mathlib/SetTheory/Ordinal/NaturalOps.lean | 7 ++-- Mathlib/SetTheory/Ordinal/Principal.lean | 39 +++++++++++++++++++---- 3 files changed, 42 insertions(+), 13 deletions(-) diff --git a/Mathlib/SetTheory/Ordinal/Arithmetic.lean b/Mathlib/SetTheory/Ordinal/Arithmetic.lean index b43f781637c4a..c8a10d44c7ce0 100644 --- a/Mathlib/SetTheory/Ordinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Ordinal/Arithmetic.lean @@ -1927,12 +1927,15 @@ theorem IsNormal.eq_iff_zero_and_succ {f g : Ordinal.{u} → Ordinal.{u}} (hf : exact H b hb⟩ /-- A two-argument version of `Ordinal.blsub`. -We don't develop a full API for this, since it's only used in a handful of existence results. -/ + +Deprecated. If you need this value explicitly, write it in terms of `iSup`. If you just want an +upper bound for the image of `op`, use that `Iio a ×ˢ Iio b` is a small set. -/ +@[deprecated (since := "2024-10-11")] def blsub₂ (o₁ o₂ : Ordinal) (op : {a : Ordinal} → (a < o₁) → {b : Ordinal} → (b < o₂) → Ordinal) : Ordinal := lsub (fun x : o₁.toType × o₂.toType => op (typein_lt_self x.1) (typein_lt_self x.2)) --- TODO: deprecate this, and replace the arguments using it by arguments about small sets. +@[deprecated (since := "2024-10-11")] theorem lt_blsub₂ {o₁ o₂ : Ordinal} (op : {a : Ordinal} → (a < o₁) → {b : Ordinal} → (b < o₂) → Ordinal) {a b : Ordinal} (ha : a < o₁) (hb : b < o₂) : op ha hb < blsub₂ o₁ o₂ op := by @@ -1942,8 +1945,6 @@ theorem lt_blsub₂ {o₁ o₂ : Ordinal} end blsub --- TODO: deprecate in favor of `sInf sᶜ`. - section mex set_option linter.deprecated false diff --git a/Mathlib/SetTheory/Ordinal/NaturalOps.lean b/Mathlib/SetTheory/Ordinal/NaturalOps.lean index 3de7f421c950a..85ff9d0c970cb 100644 --- a/Mathlib/SetTheory/Ordinal/NaturalOps.lean +++ b/Mathlib/SetTheory/Ordinal/NaturalOps.lean @@ -447,8 +447,11 @@ theorem nmul_def (a b : Ordinal) : /-- The set in the definition of `nmul` is nonempty. -/ private theorem nmul_nonempty (a b : Ordinal.{u}) : - {c : Ordinal.{u} | ∀ a' < a, ∀ b' < b, a' ⨳ b ♯ a ⨳ b' < c ♯ a' ⨳ b'}.Nonempty := - ⟨_, fun _ ha _ hb => (lt_blsub₂.{u, u, u} _ ha hb).trans_le le_self_nadd⟩ + {c : Ordinal.{u} | ∀ a' < a, ∀ b' < b, a' ⨳ b ♯ a ⨳ b' < c ♯ a' ⨳ b'}.Nonempty := by + obtain ⟨c, hc⟩ : BddAbove ((fun x ↦ x.1 ⨳ b ♯ a ⨳ x.2) '' Set.Iio a ×ˢ Set.Iio b) := + bddAbove_of_small _ + exact ⟨_, fun x hx y hy ↦ + (lt_succ_of_le <| hc <| Set.mem_image_of_mem _ <| Set.mk_mem_prod hx hy).trans_le le_self_nadd⟩ theorem nmul_nadd_lt {a' b' : Ordinal} (ha : a' < a) (hb : b' < b) : a' ⨳ b ♯ a ⨳ b' < a ⨳ b ♯ a' ⨳ b' := by diff --git a/Mathlib/SetTheory/Ordinal/Principal.lean b/Mathlib/SetTheory/Ordinal/Principal.lean index 5e32abd578e7d..411bdf03b41b2 100644 --- a/Mathlib/SetTheory/Ordinal/Principal.lean +++ b/Mathlib/SetTheory/Ordinal/Principal.lean @@ -14,7 +14,7 @@ We define principal or indecomposable ordinals, and we prove the standard proper * `Principal`: A principal or indecomposable ordinal under some binary operation. We include 0 and any other typically excluded edge cases for simplicity. -* `unbounded_principal`: Principal ordinals are unbounded. +* `not_bddAbove_principal`: Principal ordinals (under any operation) are unbounded. * `principal_add_iff_zero_or_omega0_opow`: The main characterization theorem for additive principal ordinals. * `principal_mul_iff_le_two_or_omega0_opow_opow`: The main characterization theorem for @@ -98,11 +98,35 @@ end Arbitrary /-! ### Principal ordinals are unbounded -/ -#adaptation_note /-- 2024-04-23 -After https://github.com/leanprover/lean4/pull/3965, -we need to write `lt_blsub₂.{u}` twice below, -where previously the universe annotation was not necessary. -This appears to be correct behaviour, as `lt_blsub₂.{0}` also works. -/ +/-- We give an explicit construction for a principal ordinal larger or equal than `o`. -/ +private theorem principal_nfp_iSup (op : Ordinal → Ordinal → Ordinal) (o : Ordinal) : + Principal op (nfp (fun x ↦ ⨆ y : Set.Iio x ×ˢ Set.Iio x, succ (op y.1.1 y.1.2)) o) := by + intro a b ha hb + rw [lt_nfp] at * + obtain ⟨m, ha⟩ := ha + obtain ⟨n, hb⟩ := hb + obtain h | h := le_total + ((fun x ↦ ⨆ y : Set.Iio x ×ˢ Set.Iio x, succ (op y.1.1 y.1.2))^[m] o) + ((fun x ↦ ⨆ y : Set.Iio x ×ˢ Set.Iio x, succ (op y.1.1 y.1.2))^[n] o) + · use n + 1 + rw [Function.iterate_succ'] + apply (lt_succ _).trans_le + exact Ordinal.le_iSup (fun y : Set.Iio _ ×ˢ Set.Iio _ ↦ succ (op y.1.1 y.1.2)) + ⟨_, Set.mk_mem_prod (ha.trans_le h) hb⟩ + · use m + 1 + rw [Function.iterate_succ'] + apply (lt_succ _).trans_le + exact Ordinal.le_iSup (fun y : Set.Iio _ ×ˢ Set.Iio _ ↦ succ (op y.1.1 y.1.2)) + ⟨_, Set.mk_mem_prod ha (hb.trans_le h)⟩ + +/-- Principal ordinals under any operation are unbounded. -/ +theorem not_bddAbove_principal (op : Ordinal → Ordinal → Ordinal) : + ¬ BddAbove { o | Principal op o } := by + rintro ⟨a, ha⟩ + exact ((le_nfp _ _).trans (ha (principal_nfp_iSup op (succ a)))).not_lt (lt_succ a) + +set_option linter.deprecated false in +@[deprecated (since := "2024-10-11")] theorem principal_nfp_blsub₂ (op : Ordinal → Ordinal → Ordinal) (o : Ordinal) : Principal op (nfp (fun o' => blsub₂.{u, u, u} o' o' (@fun a _ b _ => op a b)) o) := by intro a b ha hb @@ -119,7 +143,8 @@ theorem principal_nfp_blsub₂ (op : Ordinal → Ordinal → Ordinal) (o : Ordin rw [Function.iterate_succ'] exact lt_blsub₂ (@fun a _ b _ => op a b) hm (hn.trans_le h) -/-- Principal ordinals under any operation form a ZFC proper class. -/ +set_option linter.deprecated false in +@[deprecated (since := "2024-10-11")] theorem unbounded_principal (op : Ordinal → Ordinal → Ordinal) : Set.Unbounded (· < ·) { o | Principal op o } := fun o => ⟨_, principal_nfp_blsub₂ op o, (le_nfp _ o).not_lt⟩ From 797cd9177e5a503d1bdb091db11268f7138bcab0 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 21 Oct 2024 07:59:29 +0000 Subject: [PATCH 220/505] chore(ConditionalProbability): review argument implicitness (#17910) Make the proof arguments come first and make the measure argument implicit when it can be inferred from the later arguments --- Archive/Wiedijk100Theorems/BallotProblem.lean | 6 +- .../Probability/ConditionalProbability.lean | 64 +++++++++---------- Mathlib/Probability/UniformOn.lean | 26 ++++---- 3 files changed, 47 insertions(+), 49 deletions(-) diff --git a/Archive/Wiedijk100Theorems/BallotProblem.lean b/Archive/Wiedijk100Theorems/BallotProblem.lean index 3fd05ee140379..d11275c43b41a 100644 --- a/Archive/Wiedijk100Theorems/BallotProblem.lean +++ b/Archive/Wiedijk100Theorems/BallotProblem.lean @@ -216,7 +216,7 @@ theorem first_vote_pos : rw [inter_eq_right, counted_succ_succ] exact subset_union_left rw [(uniformOn_eq_zero_iff <| (countedSequence_finite _ _).image _).2 this, uniformOn, - cond_apply _ list_int_measurableSet, hint, count_injective_image List.cons_injective, + cond_apply list_int_measurableSet, hint, count_injective_image List.cons_injective, count_countedSequence, count_countedSequence, one_mul, zero_mul, add_zero, Nat.cast_add, Nat.cast_one, mul_comm, ← div_eq_mul_inv, ENNReal.div_eq_div_iff] · norm_cast @@ -270,7 +270,7 @@ theorem ballot_pos (p q : ℕ) : uniformOn (countedSequence (p + 1) (q + 1) ∩ {l | l.headI = 1}) staysPositive = uniformOn (countedSequence p (q + 1)) staysPositive := by rw [countedSequence_int_pos_counted_succ_succ, uniformOn, uniformOn, - cond_apply _ list_int_measurableSet, cond_apply _ list_int_measurableSet, + cond_apply list_int_measurableSet, cond_apply list_int_measurableSet, count_injective_image List.cons_injective] congr 1 have : (1 :: ·) '' countedSequence p (q + 1) ∩ staysPositive = @@ -297,7 +297,7 @@ theorem ballot_neg (p q : ℕ) (qp : q < p) : uniformOn (countedSequence (p + 1) (q + 1) ∩ {l | l.headI = 1}ᶜ) staysPositive = uniformOn (countedSequence (p + 1) q) staysPositive := by rw [countedSequence_int_neg_counted_succ_succ, uniformOn, uniformOn, - cond_apply _ list_int_measurableSet, cond_apply _ list_int_measurableSet, + cond_apply list_int_measurableSet, cond_apply list_int_measurableSet, count_injective_image List.cons_injective] congr 1 have : List.cons (-1) '' countedSequence (p + 1) q ∩ staysPositive = diff --git a/Mathlib/Probability/ConditionalProbability.lean b/Mathlib/Probability/ConditionalProbability.lean index aa9148f363412..64c784c227353 100644 --- a/Mathlib/Probability/ConditionalProbability.lean +++ b/Mathlib/Probability/ConditionalProbability.lean @@ -58,21 +58,18 @@ noncomputable section open ENNReal MeasureTheory MeasureTheory.Measure MeasurableSpace Set -variable {Ω Ω' α : Type*} {m : MeasurableSpace Ω} {m' : MeasurableSpace Ω'} (μ : Measure Ω) +variable {Ω Ω' α : Type*} {m : MeasurableSpace Ω} {m' : MeasurableSpace Ω'} {μ : Measure Ω} {s t : Set Ω} namespace ProbabilityTheory -section Definitions - +variable (μ) in /-- The conditional probability measure of measure `μ` on set `s` is `μ` restricted to `s` and scaled by the inverse of `μ s` (to make it a probability measure): `(μ s)⁻¹ • μ.restrict s`. -/ def cond (s : Set Ω) : Measure Ω := (μ s)⁻¹ • μ.restrict s -end Definitions - @[inherit_doc] scoped notation μ "[" s "|" t "]" => ProbabilityTheory.cond μ t s @[inherit_doc] scoped notation:max μ "[|" t "]" => ProbabilityTheory.cond μ t @@ -95,7 +92,7 @@ theorem cond_isProbabilityMeasure_of_finite (hcs : μ s ≠ 0) (hs : μ s ≠ /-- The conditional probability measure of any finite measure on any set of positive measure is a probability measure. -/ theorem cond_isProbabilityMeasure [IsFiniteMeasure μ] (hcs : μ s ≠ 0) : - IsProbabilityMeasure μ[|s] := cond_isProbabilityMeasure_of_finite μ hcs (measure_ne_top μ s) + IsProbabilityMeasure μ[|s] := cond_isProbabilityMeasure_of_finite hcs (measure_ne_top μ s) instance : IsZeroOrProbabilityMeasure μ[|s] := by constructor @@ -107,6 +104,7 @@ instance : IsZeroOrProbabilityMeasure μ[|s] := by · simp [h'] simp [ENNReal.div_self h h'] +variable (μ) in theorem cond_toMeasurable_eq : μ[|(toMeasurable μ s)] = μ[|s] := by unfold cond @@ -114,11 +112,9 @@ theorem cond_toMeasurable_eq : · simp [hnt] · simp [Measure.restrict_toMeasurable hnt] -variable {μ} in lemma cond_absolutelyContinuous : μ[|s] ≪ μ := smul_absolutelyContinuous.trans restrict_le_self.absolutelyContinuous -variable {μ} in lemma absolutelyContinuous_cond_univ [IsFiniteMeasure μ] : μ ≪ μ[|univ] := by rw [cond, restrict_univ] refine absolutelyContinuous_smul ?_ @@ -126,11 +122,11 @@ lemma absolutelyContinuous_cond_univ [IsFiniteMeasure μ] : μ ≪ μ[|univ] := section Bayes -@[simp] -theorem cond_empty : μ[|∅] = 0 := by simp [cond] +variable (μ) in +@[simp] lemma cond_empty : μ[|∅] = 0 := by simp [cond] -@[simp] -theorem cond_univ [IsProbabilityMeasure μ] : μ[|Set.univ] = μ := by +variable (μ) in +@[simp] lemma cond_univ [IsProbabilityMeasure μ] : μ[|Set.univ] = μ := by simp [cond, measure_univ, Measure.restrict_univ] @[simp] lemma cond_eq_zero (hμs : μ s ≠ ⊤) : μ[|s] = 0 ↔ μ s = 0 := by simp [cond, hμs] @@ -138,33 +134,35 @@ theorem cond_univ [IsProbabilityMeasure μ] : μ[|Set.univ] = μ := by lemma cond_eq_zero_of_meas_eq_zero (hμs : μ s = 0) : μ[|s] = 0 := by simp [hμs] /-- The axiomatic definition of conditional probability derived from a measure-theoretic one. -/ -theorem cond_apply (hms : MeasurableSet s) (t : Set Ω) : μ[t|s] = (μ s)⁻¹ * μ (s ∩ t) := by +theorem cond_apply (hms : MeasurableSet s) (μ : Measure Ω) (t : Set Ω) : + μ[t|s] = (μ s)⁻¹ * μ (s ∩ t) := by rw [cond, Measure.smul_apply, Measure.restrict_apply' hms, Set.inter_comm, smul_eq_mul] -theorem cond_apply' {t : Set Ω} (hA : MeasurableSet t) : μ[t|s] = (μ s)⁻¹ * μ (s ∩ t) := by - rw [cond, Measure.smul_apply, Measure.restrict_apply hA, Set.inter_comm, smul_eq_mul] +theorem cond_apply' (ht : MeasurableSet t) (μ : Measure Ω) : μ[t|s] = (μ s)⁻¹ * μ (s ∩ t) := by + rw [cond, Measure.smul_apply, Measure.restrict_apply ht, Set.inter_comm, smul_eq_mul] @[simp] lemma cond_apply_self (hs₀ : μ s ≠ 0) (hs : μ s ≠ ∞) : μ[s|s] = 1 := by simpa [cond] using ENNReal.inv_mul_cancel hs₀ hs -theorem cond_inter_self (hms : MeasurableSet s) (t : Set Ω) : μ[s ∩ t|s] = μ[t|s] := by - rw [cond_apply _ hms, ← Set.inter_assoc, Set.inter_self, ← cond_apply _ hms] +theorem cond_inter_self (hms : MeasurableSet s) (t : Set Ω) (μ : Measure Ω) : + μ[s ∩ t|s] = μ[t|s] := by + rw [cond_apply hms, ← Set.inter_assoc, Set.inter_self, ← cond_apply hms] theorem inter_pos_of_cond_ne_zero (hms : MeasurableSet s) (hcst : μ[t|s] ≠ 0) : 0 < μ (s ∩ t) := by refine pos_iff_ne_zero.mpr (right_ne_zero_of_mul (a := (μ s)⁻¹) ?_) convert hcst simp [hms, Set.inter_comm, cond] -theorem cond_pos_of_inter_ne_zero [IsFiniteMeasure μ] - (hms : MeasurableSet s) (hci : μ (s ∩ t) ≠ 0) : 0 < μ[|s] t := by - rw [cond_apply _ hms] +lemma cond_pos_of_inter_ne_zero [IsFiniteMeasure μ] (hms : MeasurableSet s) (hci : μ (s ∩ t) ≠ 0) : + 0 < μ[|s] t := by + rw [cond_apply hms] refine ENNReal.mul_pos ?_ hci exact ENNReal.inv_ne_zero.mpr (measure_ne_top _ _) lemma cond_cond_eq_cond_inter' (hms : MeasurableSet s) (hmt : MeasurableSet t) (hcs : μ s ≠ ∞) : μ[|s][|t] = μ[|s ∩ t] := by ext u - rw [cond_apply _ hmt, cond_apply _ hms, cond_apply _ hms, cond_apply _ (hms.inter hmt)] + rw [cond_apply hmt, cond_apply hms, cond_apply hms, cond_apply (hms.inter hmt)] obtain hst | hst := eq_or_ne (μ (s ∩ t)) 0 · have : μ (s ∩ t ∩ u) = 0 := measure_mono_null Set.inter_subset_left hst simp [this, ← Set.inter_assoc] @@ -175,30 +173,30 @@ lemma cond_cond_eq_cond_inter' (hms : MeasurableSet s) (hmt : MeasurableSet t) ( /-- Conditioning first on `s` and then on `t` results in the same measure as conditioning on `s ∩ t`. -/ -theorem cond_cond_eq_cond_inter [IsFiniteMeasure μ] (hms : MeasurableSet s) - (hmt : MeasurableSet t) : μ[|s][|t] = μ[|s ∩ t] := - cond_cond_eq_cond_inter' μ hms hmt (measure_ne_top μ s) +theorem cond_cond_eq_cond_inter (hms : MeasurableSet s) (hmt : MeasurableSet t) (μ : Measure Ω) + [IsFiniteMeasure μ] : μ[|s][|t] = μ[|s ∩ t] := + cond_cond_eq_cond_inter' hms hmt (measure_ne_top μ s) theorem cond_mul_eq_inter' (hms : MeasurableSet s) (hcs' : μ s ≠ ∞) (t : Set Ω) : μ[t|s] * μ s = μ (s ∩ t) := by obtain hcs | hcs := eq_or_ne (μ s) 0 · simp [hcs, measure_inter_null_of_null_left] - · rw [cond_apply μ hms t, mul_comm, ← mul_assoc, ENNReal.mul_inv_cancel hcs hcs', one_mul] + · rw [cond_apply hms, mul_comm, ← mul_assoc, ENNReal.mul_inv_cancel hcs hcs', one_mul] -theorem cond_mul_eq_inter [IsFiniteMeasure μ] (hms : MeasurableSet s) (t : Set Ω) : - μ[t|s] * μ s = μ (s ∩ t) := cond_mul_eq_inter' μ hms (measure_ne_top _ s) t +theorem cond_mul_eq_inter (hms : MeasurableSet s) (t : Set Ω) (μ : Measure Ω) [IsFiniteMeasure μ] : + μ[t|s] * μ s = μ (s ∩ t) := cond_mul_eq_inter' hms (measure_ne_top _ s) t /-- A version of the law of total probability. -/ -theorem cond_add_cond_compl_eq [IsFiniteMeasure μ] (hms : MeasurableSet s) : +theorem cond_add_cond_compl_eq (hms : MeasurableSet s) (μ : Measure Ω) [IsFiniteMeasure μ] : μ[t|s] * μ s + μ[t|sᶜ] * μ sᶜ = μ t := by - rw [cond_mul_eq_inter μ hms, cond_mul_eq_inter μ hms.compl, Set.inter_comm _ t, + rw [cond_mul_eq_inter hms, cond_mul_eq_inter hms.compl, Set.inter_comm _ t, Set.inter_comm _ t] exact measure_inter_add_diff t hms /-- **Bayes' Theorem** -/ -theorem cond_eq_inv_mul_cond_mul [IsFiniteMeasure μ] - (hms : MeasurableSet s) (hmt : MeasurableSet t) : μ[t|s] = (μ s)⁻¹ * μ[s|t] * μ t := by - rw [mul_assoc, cond_mul_eq_inter μ hmt s, Set.inter_comm, cond_apply _ hms] +theorem cond_eq_inv_mul_cond_mul (hms : MeasurableSet s) (hmt : MeasurableSet t) (μ : Measure Ω) + [IsFiniteMeasure μ] : μ[t|s] = (μ s)⁻¹ * μ[s|t] * μ t := by + rw [mul_assoc, cond_mul_eq_inter hmt s, Set.inter_comm, cond_apply hms] end Bayes @@ -229,7 +227,7 @@ lemma sum_meas_smul_cond_fiber {X : Ω → α} (hX : Measurable X) (μ : Measure _ = ∑ x, μ (X ⁻¹' {x} ∩ E) := by simp only [Measure.coe_finset_sum, Measure.coe_smul, Finset.sum_apply, Pi.smul_apply, smul_eq_mul] - simp_rw [mul_comm (μ _), cond_mul_eq_inter _ (hX (.singleton _))] + simp_rw [mul_comm (μ _), cond_mul_eq_inter (hX (.singleton _))] _ = _ := by have : ⋃ x ∈ Finset.univ, X ⁻¹' {x} ∩ E = E := by ext; simp rw [← measure_biUnion_finset _ fun _ _ ↦ (hX (.singleton _)).inter hE, this] diff --git a/Mathlib/Probability/UniformOn.lean b/Mathlib/Probability/UniformOn.lean index 8b1de52a5055d..524e33e17554c 100644 --- a/Mathlib/Probability/UniformOn.lean +++ b/Mathlib/Probability/UniformOn.lean @@ -79,7 +79,7 @@ alias finite_of_condCount_ne_zero := finite_of_uniformOn_ne_zero theorem uniformOn_univ [Fintype Ω] {s : Set Ω} : uniformOn Set.univ s = Measure.count s / Fintype.card Ω := by - rw [uniformOn, cond_apply _ MeasurableSet.univ, ← ENNReal.div_eq_inv_mul, Set.univ_inter] + rw [uniformOn, cond_apply MeasurableSet.univ, ← ENNReal.div_eq_inv_mul, Set.univ_inter] congr rw [← Finset.coe_univ, Measure.count_apply, Finset.univ.tsum_subtype' fun _ => (1 : ENNReal)] · simp [Finset.card_univ] @@ -101,7 +101,7 @@ alias condCount_isProbabilityMeasure := uniformOn_isProbabilityMeasure theorem uniformOn_singleton (ω : Ω) (t : Set Ω) [Decidable (ω ∈ t)] : uniformOn {ω} t = if ω ∈ t then 1 else 0 := by - rw [uniformOn, cond_apply _ (measurableSet_singleton ω), Measure.count_singleton, inv_one, + rw [uniformOn, cond_apply (measurableSet_singleton ω), Measure.count_singleton, inv_one, one_mul] split_ifs · rw [(by simpa : ({ω} : Set Ω) ∩ t = {ω}), Measure.count_singleton] @@ -113,13 +113,13 @@ alias condCount_singleton := uniformOn_singleton variable {s t u : Set Ω} theorem uniformOn_inter_self (hs : s.Finite) : uniformOn s (s ∩ t) = uniformOn s t := by - rw [uniformOn, cond_inter_self _ hs.measurableSet] + rw [uniformOn, cond_inter_self hs.measurableSet] @[deprecated (since := "2024-10-09")] alias condCount_inter_self := uniformOn_inter_self theorem uniformOn_self (hs : s.Finite) (hs' : s.Nonempty) : uniformOn s s = 1 := by - rw [uniformOn, cond_apply _ hs.measurableSet, Set.inter_self, ENNReal.inv_mul_cancel] + rw [uniformOn, cond_apply hs.measurableSet, Set.inter_self, ENNReal.inv_mul_cancel] · exact fun h => hs'.ne_empty <| Measure.empty_of_count_eq_zero h · exact (Measure.count_apply_lt_top.2 hs).ne @@ -138,7 +138,7 @@ alias condCount_eq_one_of := uniformOn_eq_one_of theorem pred_true_of_uniformOn_eq_one (h : uniformOn s t = 1) : s ⊆ t := by have hsf := finite_of_uniformOn_ne_zero (by rw [h]; exact one_ne_zero) - rw [uniformOn, cond_apply _ hsf.measurableSet, mul_comm] at h + rw [uniformOn, cond_apply hsf.measurableSet, mul_comm] at h replace h := ENNReal.eq_inv_of_mul_eq_one_left h rw [inv_inv, Measure.count_apply_finite _ hsf, Measure.count_apply_finite _ (hsf.inter_of_left _), Nat.cast_inj] at h @@ -150,7 +150,7 @@ theorem pred_true_of_uniformOn_eq_one (h : uniformOn s t = 1) : s ⊆ t := by alias pred_true_of_condCount_eq_one := pred_true_of_uniformOn_eq_one theorem uniformOn_eq_zero_iff (hs : s.Finite) : uniformOn s t = 0 ↔ s ∩ t = ∅ := by - simp [uniformOn, cond_apply _ hs.measurableSet, Measure.count_apply_eq_top, Set.not_infinite.2 hs, + simp [uniformOn, cond_apply hs.measurableSet, Measure.count_apply_eq_top, Set.not_infinite.2 hs, Measure.count_apply_finite _ (hs.inter_of_left _)] @[deprecated (since := "2024-10-09")] @@ -167,8 +167,8 @@ theorem uniformOn_inter (hs : s.Finite) : by_cases hst : s ∩ t = ∅ · rw [hst, uniformOn_empty_meas, Measure.coe_zero, Pi.zero_apply, zero_mul, uniformOn_eq_zero_iff hs, ← Set.inter_assoc, hst, Set.empty_inter] - rw [uniformOn, uniformOn, cond_apply _ hs.measurableSet, cond_apply _ hs.measurableSet, - cond_apply _ (hs.inter_of_left _).measurableSet, mul_comm _ (Measure.count (s ∩ t)), + rw [uniformOn, uniformOn, cond_apply hs.measurableSet, cond_apply hs.measurableSet, + cond_apply (hs.inter_of_left _).measurableSet, mul_comm _ (Measure.count (s ∩ t)), ← mul_assoc, mul_comm _ (Measure.count (s ∩ t)), ← mul_assoc, ENNReal.mul_inv_cancel, one_mul, mul_comm, Set.inter_assoc] · rwa [← Measure.count_eq_zero_iff] at hst @@ -187,8 +187,8 @@ alias condCount_inter' := uniformOn_inter' theorem uniformOn_union (hs : s.Finite) (htu : Disjoint t u) : uniformOn s (t ∪ u) = uniformOn s t + uniformOn s u := by - rw [uniformOn, cond_apply _ hs.measurableSet, cond_apply _ hs.measurableSet, - cond_apply _ hs.measurableSet, Set.inter_union_distrib_left, measure_union, mul_add] + rw [uniformOn, cond_apply hs.measurableSet, cond_apply hs.measurableSet, + cond_apply hs.measurableSet, Set.inter_union_distrib_left, measure_union, mul_add] exacts [htu.mono inf_le_right inf_le_right, (hs.inter_of_left _).measurableSet] @[deprecated (since := "2024-10-09")] @@ -209,9 +209,9 @@ theorem uniformOn_disjoint_union (hs : s.Finite) (ht : t.Finite) (hst : Disjoint · simp · simp [uniformOn_self ht ht'] · simp [uniformOn_self hs hs'] - rw [uniformOn, uniformOn, uniformOn, cond_apply _ hs.measurableSet, - cond_apply _ ht.measurableSet, cond_apply _ (hs.union ht).measurableSet, - cond_apply _ (hs.union ht).measurableSet, cond_apply _ (hs.union ht).measurableSet] + rw [uniformOn, uniformOn, uniformOn, cond_apply hs.measurableSet, + cond_apply ht.measurableSet, cond_apply (hs.union ht).measurableSet, + cond_apply (hs.union ht).measurableSet, cond_apply (hs.union ht).measurableSet] conv_lhs => rw [Set.union_inter_cancel_left, Set.union_inter_cancel_right, mul_comm (Measure.count (s ∪ t))⁻¹, mul_comm (Measure.count (s ∪ t))⁻¹, ← mul_assoc, From 1fad8e469361789e5a365c1ab766336daa17b2b0 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 21 Oct 2024 08:12:54 +0000 Subject: [PATCH 221/505] feat: the pi sigma-algebra is generated by projections (#17828) This is basically the definition, but it's not quite immediate. From GibbsMeasure --- Mathlib/MeasureTheory/Constructions/Pi.lean | 3 +++ Mathlib/MeasureTheory/MeasurableSpace/Basic.lean | 3 +++ 2 files changed, 6 insertions(+) diff --git a/Mathlib/MeasureTheory/Constructions/Pi.lean b/Mathlib/MeasureTheory/Constructions/Pi.lean index 40341ab6df349..2a10e3da58fe6 100644 --- a/Mathlib/MeasureTheory/Constructions/Pi.lean +++ b/Mathlib/MeasureTheory/Constructions/Pi.lean @@ -62,6 +62,9 @@ variable {ι ι' : Type*} {α : ι → Type*} /-! We start with some measurability properties -/ +lemma MeasurableSpace.pi_eq_generateFrom_projections {mα : ∀ i, MeasurableSpace (α i)} : + pi = generateFrom {B | ∃ (i : ι) (A : Set (α i)), MeasurableSet A ∧ eval i ⁻¹' A = B} := by + simp only [pi, ← generateFrom_iUnion_measurableSet, iUnion_setOf, measurableSet_comap] /-- Boxes formed by π-systems form a π-system. -/ theorem IsPiSystem.pi {C : ∀ i, Set (Set (α i))} (hC : ∀ i, IsPiSystem (C i)) : diff --git a/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean b/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean index 49c7a06029ead..b252ea984d167 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean @@ -90,6 +90,9 @@ protected def comap (f : α → β) (m : MeasurableSpace β) : MeasurableSpace let ⟨s', hs'⟩ := Classical.axiom_of_choice hs ⟨⋃ i, s' i, m.measurableSet_iUnion _ fun i => (hs' i).left, by simp [hs']⟩ +lemma measurableSet_comap {m : MeasurableSpace β} : + MeasurableSet[m.comap f] s ↔ ∃ s', MeasurableSet[m] s' ∧ f ⁻¹' s' = s := .rfl + theorem comap_eq_generateFrom (m : MeasurableSpace β) (f : α → β) : m.comap f = generateFrom { t | ∃ s, MeasurableSet s ∧ f ⁻¹' s = t } := (@generateFrom_measurableSet _ (.comap f m)).symm From 1404d2bd157a39d32b67ba6655cccb6c4d50e799 Mon Sep 17 00:00:00 2001 From: D-Thomine <100795491+D-Thomine@users.noreply.github.com> Date: Mon, 21 Oct 2024 08:28:17 +0000 Subject: [PATCH 222/505] feat(ENNReal): liminf/limsup/iInf/iSup and multiplication (#17656) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR adds a few lemmas about multiplication and liminf/limsup/iInf/iSup: `mul_iInf_le_iInf_mul`, `iSup_mul_le_mul_iSup`, `mul_liminf_le_liminf_mul` and three avatars of the latter. The strategy adopted is the same as for similar lemmas in EReal (`EReal.add_iInf_le_iInf_add`...): we suffer to prove the key result `mul_le_of_forall_mul_le`, and everything is deduced from there. Note that `ENNReal.mul_le_of_forall_mul_le` could be deduced much faster from `EReal.add_le_of_forall_add_le` using the exp (both as an order isomorphism and as a "group" morphism). However, I wanted to avoid importing special functions, which unfortunately leads to a much longer and painful proof. This also generalizes some lemmas in #15373 : `limsup_mul_le'` therein can be replaced by `limsup_mul_le_mul_limsup` with some easy plug-and-play; if `_root_.Real.limsup_mul_le` is still needed, I can write another PR to deal with it (most of the work is already done anyways -- I think I can prove `mul_le_of_forall_mul_le` in NNReal, and leverage it to simplify the proof in ENNReal). Co-authored-by: Yaël Dillies --- Mathlib/Data/ENNReal/Inv.lean | 29 +++++++++++++++++++++++++ Mathlib/Order/Filter/ENNReal.lean | 1 + Mathlib/Topology/Instances/ENNReal.lean | 27 +++++++++++++++++++++++ 3 files changed, 57 insertions(+) diff --git a/Mathlib/Data/ENNReal/Inv.lean b/Mathlib/Data/ENNReal/Inv.lean index c0a3bde921b26..b9f1884fa6b64 100644 --- a/Mathlib/Data/ENNReal/Inv.lean +++ b/Mathlib/Data/ENNReal/Inv.lean @@ -444,6 +444,27 @@ theorem sub_half (h : a ≠ ∞) : a - a / 2 = a / 2 := ENNReal.sub_eq_of_eq_add theorem one_sub_inv_two : (1 : ℝ≥0∞) - 2⁻¹ = 2⁻¹ := by simpa only [div_eq_mul_inv, one_mul] using sub_half one_ne_top +private lemma exists_lt_mul_left {a b c : ℝ≥0∞} (hc : c < a * b) : ∃ a' < a, c < a' * b := by + obtain ⟨a', hc, ha'⟩ := exists_between (ENNReal.div_lt_of_lt_mul hc) + exact ⟨_, ha', (ENNReal.div_lt_iff (.inl <| by rintro rfl; simp at *) + (.inr <| by rintro rfl; simp at *)).1 hc⟩ + +private lemma exists_lt_mul_right {a b c : ℝ≥0∞} (hc : c < a * b) : ∃ b' < b, c < a * b' := by + simp_rw [mul_comm a] at hc ⊢; exact exists_lt_mul_left hc + +lemma mul_le_of_forall_lt {a b c : ℝ≥0∞} (h : ∀ a' < a, ∀ b' < b, a' * b' ≤ c) : a * b ≤ c := by + refine le_of_forall_ge_of_dense fun d hd ↦ ?_ + obtain ⟨a', ha', hd⟩ := exists_lt_mul_left hd + obtain ⟨b', hb', hd⟩ := exists_lt_mul_right hd + exact le_trans hd.le <| h _ ha' _ hb' + +lemma le_mul_of_forall_lt {a b c : ℝ≥0∞} (h₁ : a ≠ 0 ∨ b ≠ ∞) (h₂ : a ≠ ∞ ∨ b ≠ 0) + (h : ∀ a' > a, ∀ b' > b, c ≤ a' * b') : c ≤ a * b := by + rw [← ENNReal.inv_le_inv, ENNReal.mul_inv h₁ h₂] + exact mul_le_of_forall_lt fun a' ha' b' hb' ↦ ENNReal.le_inv_iff_le_inv.1 <| + (h _ (ENNReal.lt_inv_iff_lt_inv.1 ha') _ (ENNReal.lt_inv_iff_lt_inv.1 hb')).trans_eq + (ENNReal.mul_inv (Or.inr hb'.ne_top) (Or.inl ha'.ne_top)).symm + /-- The birational order isomorphism between `ℝ≥0∞` and the unit interval `Set.Iic (1 : ℝ≥0∞)`. -/ @[simps! apply_coe] def orderIsoIicOneBirational : ℝ≥0∞ ≃o Iic (1 : ℝ≥0∞) := by @@ -731,6 +752,14 @@ lemma inv_iSup (f : ι → ℝ≥0∞) : (⨆ i, f i)⁻¹ = ⨅ i, (f i)⁻¹ : lemma inv_sInf (s : Set ℝ≥0∞) : (sInf s)⁻¹ = ⨆ a ∈ s, a⁻¹ := by simp [sInf_eq_iInf, inv_iInf] lemma inv_sSup (s : Set ℝ≥0∞) : (sSup s)⁻¹ = ⨅ a ∈ s, a⁻¹ := by simp [sSup_eq_iSup, inv_iSup] +lemma le_iInf_mul {ι : Type*} (u v : ι → ℝ≥0∞) : + (⨅ i, u i) * ⨅ i, v i ≤ ⨅ i, u i * v i := + le_iInf fun i ↦ mul_le_mul' (iInf_le u i) (iInf_le v i) + +lemma iSup_mul_le {ι : Type*} {u v : ι → ℝ≥0∞} : + ⨆ i, u i * v i ≤ (⨆ i, u i) * ⨆ i, v i := + iSup_le fun i ↦ mul_le_mul' (le_iSup u i) (le_iSup v i) + lemma add_iSup [Nonempty ι] (f : ι → ℝ≥0∞) : a + ⨆ i, f i = ⨆ i, a + f i := by obtain rfl | ha := eq_or_ne a ∞ · simp diff --git a/Mathlib/Order/Filter/ENNReal.lean b/Mathlib/Order/Filter/ENNReal.lean index 8acfdf4fae082..e6578ed67d021 100644 --- a/Mathlib/Order/Filter/ENNReal.lean +++ b/Mathlib/Order/Filter/ENNReal.lean @@ -62,6 +62,7 @@ theorem limsup_const_mul [CountableInterFilter f] {u : α → ℝ≥0∞} {a : have hfu : f.limsup u ≠ 0 := mt limsup_eq_zero_iff.1 hu simp only [ha_top, top_mul', h_top_le, hfu, ite_false] +/-- See also `limsup_mul_le'`.-/ theorem limsup_mul_le [CountableInterFilter f] (u v : α → ℝ≥0∞) : f.limsup (u * v) ≤ f.limsup u * f.limsup v := calc diff --git a/Mathlib/Topology/Instances/ENNReal.lean b/Mathlib/Topology/Instances/ENNReal.lean index f5c72e34f691d..bee9dc9f1a492 100644 --- a/Mathlib/Topology/Instances/ENNReal.lean +++ b/Mathlib/Topology/Instances/ENNReal.lean @@ -1364,6 +1364,33 @@ lemma liminf_const_sub (F : Filter ι) [NeBot F] (f : ι → ℝ≥0∞) {c : (Antitone.map_limsSup_of_continuousAt (F := F.map f) (f := fun (x : ℝ≥0∞) ↦ c - x) (fun _ _ h ↦ tsub_le_tsub_left h c) (continuous_sub_left c_ne_top).continuousAt).symm +lemma le_limsup_mul {α : Type*} {f : Filter α} {u v : α → ℝ≥0∞} : + limsup u f * liminf v f ≤ limsup (u * v) f := + mul_le_of_forall_lt fun a a_u b b_v ↦ (le_limsup_iff).2 fun c c_ab ↦ + Frequently.mono (Frequently.and_eventually ((frequently_lt_of_lt_limsup) a_u) + ((eventually_lt_of_lt_liminf) b_v)) fun _ ab_x ↦ c_ab.trans (mul_lt_mul ab_x.1 ab_x.2) + +/-- See also `ENNReal.limsup_mul_le`.-/ +lemma limsup_mul_le' {α : Type*} {f : Filter α} {u v : α → ℝ≥0∞} + (h : limsup u f ≠ 0 ∨ limsup v f ≠ ∞) (h' : limsup u f ≠ ∞ ∨ limsup v f ≠ 0) : + limsup (u * v) f ≤ limsup u f * limsup v f := by + refine le_mul_of_forall_lt h h' fun a a_u b b_v ↦ (limsup_le_iff).2 fun c c_ab ↦ ?_ + filter_upwards [eventually_lt_of_limsup_lt a_u, eventually_lt_of_limsup_lt b_v] with x a_x b_x + exact (mul_lt_mul a_x b_x).trans c_ab + +lemma le_liminf_mul {α : Type*} {f : Filter α} {u v : α → ℝ≥0∞} : + liminf u f * liminf v f ≤ liminf (u * v) f := by + refine mul_le_of_forall_lt fun a a_u b b_v ↦ (le_liminf_iff).2 fun c c_ab ↦ ?_ + filter_upwards [eventually_lt_of_lt_liminf a_u, eventually_lt_of_lt_liminf b_v] with x a_x b_x + exact c_ab.trans (mul_lt_mul a_x b_x) + +lemma liminf_mul_le {α : Type*} {f : Filter α} {u v : α → ℝ≥0∞} + (h : limsup u f ≠ 0 ∨ liminf v f ≠ ∞) (h' : limsup u f ≠ ∞ ∨ liminf v f ≠ 0) : + liminf (u * v) f ≤ limsup u f * liminf v f := + le_mul_of_forall_lt h h' fun a a_u b b_v ↦ (liminf_le_iff).2 fun c c_ab ↦ + Frequently.mono (((frequently_lt_of_liminf_lt) b_v).and_eventually + ((eventually_lt_of_limsup_lt) a_u)) fun _ ab_x ↦ (mul_lt_mul ab_x.2 ab_x.1).trans c_ab + /-- If `xs : ι → ℝ≥0∞` is bounded, then we have `liminf (toReal ∘ xs) = toReal (liminf xs)`. -/ lemma liminf_toReal_eq {ι : Type*} {F : Filter ι} [NeBot F] {b : ℝ≥0∞} (b_ne_top : b ≠ ∞) {xs : ι → ℝ≥0∞} (le_b : ∀ᶠ i in F, xs i ≤ b) : From 3b842b0ccb16e851cbdd3948cf8c2cb1e513125b Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 21 Oct 2024 08:47:12 +0000 Subject: [PATCH 223/505] chore: cleanup adaptation notes from 2024-09-06 (#17990) --- Mathlib/Data/List/Forall2.lean | 8 ++------ Mathlib/Data/Nat/Defs.lean | 7 +------ Mathlib/NumberTheory/LegendreSymbol/JacobiSymbol.lean | 8 +------- Mathlib/SetTheory/Cardinal/Finite.lean | 6 +----- 4 files changed, 5 insertions(+), 24 deletions(-) diff --git a/Mathlib/Data/List/Forall2.lean b/Mathlib/Data/List/Forall2.lean index 20808976c60ab..79f5bf6c1cb0a 100644 --- a/Mathlib/Data/List/Forall2.lean +++ b/Mathlib/Data/List/Forall2.lean @@ -84,18 +84,14 @@ theorem forall₂_cons_right_iff {b l u} : match u, h with | _, ⟨_, _, h₁, h₂, rfl⟩ => Forall₂.cons h₁ h₂ -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefixes below. --/ theorem forall₂_and_left {p : α → Prop} : ∀ l u, Forall₂ (fun a b => p a ∧ R a b) l u ↔ (∀ a ∈ l, p a) ∧ Forall₂ R l u | [], u => by simp only [forall₂_nil_left_iff, forall_prop_of_false (not_mem_nil _), imp_true_iff, true_and] | a :: l, u => by - simp only [forall₂_and_left l, forall₂_cons_left_iff, forall_mem_cons, _root_.and_assoc, + simp only [forall₂_and_left l, forall₂_cons_left_iff, forall_mem_cons, and_assoc, @and_comm _ (p a), @and_left_comm _ (p a), exists_and_left] - simp only [_root_.and_comm, _root_.and_assoc, and_left_comm, ← exists_and_right] + simp only [and_comm, and_assoc, and_left_comm, ← exists_and_right] @[simp] theorem forall₂_map_left_iff {f : γ → α} : diff --git a/Mathlib/Data/Nat/Defs.lean b/Mathlib/Data/Nat/Defs.lean index 4c0204101ae2e..2d0223819be8a 100644 --- a/Mathlib/Data/Nat/Defs.lean +++ b/Mathlib/Data/Nat/Defs.lean @@ -165,13 +165,8 @@ lemma pred_eq_of_eq_succ (H : m = n.succ) : m.pred = n := by simp [H] @[simp] lemma pred_eq_succ_iff : n - 1 = m + 1 ↔ n = m + 2 := by cases n <;> constructor <;> rintro ⟨⟩ <;> rfl -#adaptation_note -/-- -After nightly-2024-09-06 we can remove both the `_root_` prefixes below. --/ lemma forall_lt_succ : (∀ m < n + 1, p m) ↔ (∀ m < n, p m) ∧ p n := by - simp only [Nat.lt_succ_iff, Nat.le_iff_lt_or_eq, _root_.or_comm, forall_eq_or_imp, - _root_.and_comm] + simp only [Nat.lt_succ_iff, Nat.le_iff_lt_or_eq, or_comm, forall_eq_or_imp, and_comm] lemma exists_lt_succ : (∃ m < n + 1, p m) ↔ (∃ m < n, p m) ∨ p n := by rw [← not_iff_not] diff --git a/Mathlib/NumberTheory/LegendreSymbol/JacobiSymbol.lean b/Mathlib/NumberTheory/LegendreSymbol/JacobiSymbol.lean index a331c2298bfd9..f408526e3fb67 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/JacobiSymbol.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/JacobiSymbol.lean @@ -156,20 +156,14 @@ theorem mul_left (a₁ a₂ : ℤ) (b : ℕ) : J(a₁ * a₂ | b) = J(a₁ | b) (f := fun x ↦ @legendreSym x {out := prime_of_mem_primeFactorsList x.2} a₁) (g := fun x ↦ @legendreSym x {out := prime_of_mem_primeFactorsList x.2} a₂) -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefixes below. --/ /-- The symbol `J(a | b)` vanishes iff `a` and `b` are not coprime (assuming `b ≠ 0`). -/ theorem eq_zero_iff_not_coprime {a : ℤ} {b : ℕ} [NeZero b] : J(a | b) = 0 ↔ a.gcd b ≠ 1 := List.prod_eq_zero_iff.trans (by rw [List.mem_pmap, Int.gcd_eq_natAbs, Ne, Prime.not_coprime_iff_dvd] - -- Porting note: Initially, `and_assoc'` and `and_comm'` were used on line 164 but they have - -- been deprecated so we replace them with `and_assoc` and `and_comm` simp_rw [legendreSym.eq_zero_iff _ _, intCast_zmod_eq_zero_iff_dvd, mem_primeFactorsList (NeZero.ne b), ← Int.natCast_dvd, Int.natCast_dvd_natCast, exists_prop, - _root_.and_assoc, _root_.and_comm]) + and_assoc, _root_.and_comm]) /-- The symbol `J(a | b)` is nonzero when `a` and `b` are coprime. -/ protected theorem ne_zero {a : ℤ} {b : ℕ} (h : a.gcd b = 1) : J(a | b) ≠ 0 := by diff --git a/Mathlib/SetTheory/Cardinal/Finite.lean b/Mathlib/SetTheory/Cardinal/Finite.lean index 54b12ce1da78f..350d46ef5b8f2 100644 --- a/Mathlib/SetTheory/Cardinal/Finite.lean +++ b/Mathlib/SetTheory/Cardinal/Finite.lean @@ -95,10 +95,6 @@ protected theorem bijective_iff_injective_and_card [Finite β] (f : α → β) : rw [← and_congr_right_iff, ← Bijective, card_eq_fintype_card, card_eq_fintype_card, Fintype.bijective_iff_injective_and_card] -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefixes below. --/ protected theorem bijective_iff_surjective_and_card [Finite α] (f : α → β) : Bijective f ↔ Surjective f ∧ Nat.card α = Nat.card β := by classical @@ -107,7 +103,7 @@ protected theorem bijective_iff_surjective_and_card [Finite α] (f : α → β) have := Fintype.ofFinite α have := Fintype.ofSurjective f h revert h - rw [← and_congr_left_iff, ← Bijective, ← _root_.and_comm, + rw [← and_congr_left_iff, ← Bijective, ← and_comm, card_eq_fintype_card, card_eq_fintype_card, Fintype.bijective_iff_surjective_and_card] theorem _root_.Function.Injective.bijective_of_nat_card_le [Finite β] {f : α → β} From 5958e95eebccc498ea68b5a74d00bd9d2ce5804c Mon Sep 17 00:00:00 2001 From: grunweg Date: Mon, 21 Oct 2024 08:47:13 +0000 Subject: [PATCH 224/505] chore: fix some nits from #12339 (#17994) --- Mathlib/Tactic/Linter/Multigoal.lean | 4 ++-- test/Multigoal.lean | 6 +++--- 2 files changed, 5 insertions(+), 5 deletions(-) diff --git a/Mathlib/Tactic/Linter/Multigoal.lean b/Mathlib/Tactic/Linter/Multigoal.lean index 6d459f2968179..5502f129e1887 100644 --- a/Mathlib/Tactic/Linter/Multigoal.lean +++ b/Mathlib/Tactic/Linter/Multigoal.lean @@ -13,7 +13,7 @@ There is an exception: a tactic that closes *all* remaining goals is allowed. There are a few tactics, such as `skip`, `swap` or the `try` combinator, that are intended to work specifically in such a situation. -Otherwise, the mathlib style guide ask that goals be be focused until there is only one active goal +Otherwise, the mathlib style guide ask that goals be focused until there is only one active goal at a time. If this focusing does not happen, the linter emits a warning. Typically, the focusing is achieved by the `cdot`: `·`, but, e.g., `focus` or `on_goal x` @@ -126,7 +126,7 @@ which * leave at least one goal that was not present before it ran; * are not excluded through `exclusions` or `ignoreBranch`; -together with the total number of goals +together with the total number of goals. -/ partial def getManyGoals : InfoTree → Array (Syntax × Nat) diff --git a/test/Multigoal.lean b/test/Multigoal.lean index 068e62456600e..a18e6bf88aad5 100644 --- a/test/Multigoal.lean +++ b/test/Multigoal.lean @@ -93,7 +93,7 @@ example (p : Bool) : 0 = 0 := by case' true => rfl #guard_msgs in --- `assumption'` is allowed, as it is useful precisely when there are multiple active goals +-- `assumption'` is allowed, as it is useful precisely when there are multiple active goals. example (p : Bool) (f : False) {h : 0 = 0} : 0 = 0 ∧ 0 = 1 := by cases p <;> constructor @@ -101,7 +101,7 @@ example (p : Bool) (f : False) {h : 0 = 0} : 0 = 0 ∧ 0 = 1 := by any_goals cases f #guard_msgs in --- `focus` is ignored +-- `focus` is ignored. example : True ∧ True := by constructor focus @@ -114,7 +114,7 @@ example : 1 = 1 := by sleep_heartbeats 1000 rfl --- we test that a tactic closing all remaining goals does not trigger the linter +-- We test that a tactic closing all remaining goals does not trigger the linter. macro "bi_trivial" : tactic => `(tactic| (trivial; trivial)) set_option linter.style.multiGoal true in From 26449dce1eb242b71e95addf4ae8be74f0745567 Mon Sep 17 00:00:00 2001 From: Fabrizio Barroero Date: Mon, 21 Oct 2024 09:28:57 +0000 Subject: [PATCH 225/505] feat(NumberTheory/Ostrowski): Ostrowski's Theorem for the rationals (#17138) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We complete the proof of the Archimedean case of Ostrowski's Theorem for the rational numbers: ```lean theorem mulRingNorm_equiv_standard_of_unbounded (notbdd: ¬ ∀ (n : ℕ), f n ≤ 1) : MulRingNorm.equiv f mulRingNorm_real ``` This concludes the proof of the theorem: ```lean theorem mulRingNorm_equiv_standard_or_padic (f : MulRingNorm ℚ) (hf_nontriv : f ≠ 1) : (MulRingNorm.equiv f mulRingNorm_real) ∨ ∃! p, ∃ (hp : Fact (p.Prime)), MulRingNorm.equiv f (mulRingNorm_padic p) ``` Co-authored-by: David Kurniadi Angdinata [dka31@cantab.ac.uk](mailto:dka31@cantab.ac.uk) Co-authored-by: Laura Capuano [laura.capuano@uniroma3.it](mailto:laura.capuano@uniroma3.it) Co-authored-by: Nirvana Coppola [nirvanac93@gmail.com](mailto:nirvanac93@gmail.com) Co-authored-by: María Inés de Frutos Fernández [maria.defrutos@uam.es](mailto:maria.defrutos@uam.es) Co-authored-by: Sam van Gool [vangool@irif.fr](mailto:vangool@irif.fr) Co-authored-by: Silvain Rideau-Kikuchi [silvain.rideau-kikuchi@ens.fr](mailto:silvain.rideau-kikuchi@ens.fr) Co-authored-by: Amos Turchet [amos.turchet@uniroma3.it](mailto:amos.turchet@uniroma3.it) Co-authored-by: Francesco Veneziano [veneziano@dima.unige.it](mailto:veneziano@dima.unige.it) Co-authored-by: Yaël Dillies Co-authored-by: Amos Co-authored-by: amosturchet --- Mathlib/NumberTheory/Ostrowski.lean | 173 +++++++++++++++++++++++++++- 1 file changed, 172 insertions(+), 1 deletion(-) diff --git a/Mathlib/NumberTheory/Ostrowski.lean b/Mathlib/NumberTheory/Ostrowski.lean index 03a573e8490f5..98af768e8637a 100644 --- a/Mathlib/NumberTheory/Ostrowski.lean +++ b/Mathlib/NumberTheory/Ostrowski.lean @@ -19,6 +19,12 @@ import Mathlib.NumberTheory.Padics.PadicNorm Ostrowski's Theorem for the field `ℚ`: every absolute value on `ℚ` is equivalent to either a `p`-adic absolute value or to the standard Archimedean (Euclidean) absolute value. +## Main results + +- `mulRingNorm_equiv_standard_or_padic`: given an absolute value on `ℚ`, it is equivalent to the +standard Archimedean (Euclidean) absolute value or to a `p`-adic absolute value for some prime +number `p`. + ## TODO Extend to arbitrary number fields. @@ -34,7 +40,8 @@ Extend to arbitrary number fields. ring norm, ostrowski -/ -/- ## Preliminary lemmas on limits -/ +/- ## Preliminary lemmas on limits and lists -/ + open Filter Nat Real Topology @@ -55,6 +62,29 @@ private lemma tendsto_nat_rpow_div : simp_rw [← one_div] exact tendsto_rpow_div +-- Multiplication by a constant moves in a List.sum +private lemma list_mul_sum {R : Type*} [CommSemiring R] {T : Type*} (l : List T) (y : R) : ∀ x : R, + List.sum (List.mapIdx (fun i _ => x * y ^ i) (l)) = + x * List.sum (List.mapIdx (fun i _ => y ^ i) (l)) := by + induction l with + | nil => simp only [List.mapIdx_nil, List.sum_nil, mul_zero, forall_const] + | cons head tail ih => + intro x + simp_rw [List.mapIdx_cons, pow_zero, mul_one, List.sum_cons, mul_add, mul_one] + have (a : ℕ) : y ^ (a + 1) = y * y ^ a := by ring + simp_rw [this, ← mul_assoc, ih, ← mul_assoc] + +-- Geometric sum for lists +private lemma list_geom {T : Type*} {F : Type*} [Field F] (l : List T) {y : F} (hy : y ≠ 1) : + List.sum (List.mapIdx (fun i _ => y ^ i) l) = (y ^ l.length - 1) / (y - 1) := by + induction l with + | nil => simp only [List.mapIdx_nil, List.sum_nil, List.length_nil, pow_zero, sub_self, zero_div] + | cons head tail ih => + simp only [List.mapIdx_cons, pow_zero, List.sum_cons, List.length_cons] + have (a : ℕ ) : y ^ (a + 1) = y * y ^ a := by ring + simp_rw [this, list_mul_sum, ih] + simp only [mul_div, ← same_add_div (sub_ne_zero.2 hy), mul_sub, mul_one, sub_add_sub_cancel'] + namespace Rat.MulRingNorm open Int @@ -261,6 +291,23 @@ section Archimedean /-! ## Archimedean case -/ +/-- The usual absolute value on ℚ. -/ +def mulRingNorm_real : MulRingNorm ℚ := +{ toFun := fun x : ℚ ↦ |x| + map_zero' := by simp only [Rat.cast_zero, abs_zero] + add_le' := norm_add_le + neg' := norm_neg + eq_zero_of_map_eq_zero' := by simp only [abs_eq_zero, Rat.cast_eq_zero, imp_self, forall_const] + map_one' := by simp only [Rat.cast_one, abs_one] + map_mul' := by + simp only [Rat.cast_mul] + exact_mod_cast abs_mul +} + +@[simp] lemma mul_ring_norm_eq_abs (r : ℚ) : mulRingNorm_real r = |r| := by + simp only [Rat.cast_abs] + rfl + /-! ## Preliminary result -/ /-- Given an two integers `n, m` with `m > 1` the mulRingNorm of `n` is bounded by @@ -344,6 +391,130 @@ lemma one_lt_of_not_bounded (notbdd : ¬ ∀ n : ℕ, f n ≤ 1) {n₀ : ℕ} (h tendsto_root_atTop_nhds_one (by positivity) exact hnlim.mul tendsto_nat_rpow_div +/-! ## step 2: given m,n ≥ 2 and |m|=m^s, |n|=n^t for s,t > 0, we have t ≤ s -/ + +variable {m n : ℕ} (hm : 1 < m) (hn : 1 < n) (notbdd : ¬ ∀ (n : ℕ), f n ≤ 1) + +include hm notbdd in +private lemma expr_pos : 0 < m * f m / (f m - 1) := by + apply div_pos (mul_pos (mod_cast zero_lt_of_lt hm) + (map_pos_of_ne_zero f (mod_cast ne_zero_of_lt hm))) + linarith only [one_lt_of_not_bounded notbdd hm] + +include hn hm notbdd in +private lemma param_upperbound {k : ℕ} (hk : k ≠ 0) : + f n ≤ (m * f m / (f m - 1)) ^ (k : ℝ)⁻¹ * (f m) ^ (logb m n) := by + have h_ineq1 {m n : ℕ} (hm : 1 < m) (hn : 1 < n) : + f n ≤ (m * f m / (f m - 1)) * (f m) ^ (logb m n) := by + let d := Nat.log m n + calc + f n ≤ ((Nat.digits m n).mapIdx fun i _ ↦ m * f m ^ i).sum := + mulRingNorm_apply_le_sum_digits n hm + _ = m * ((Nat.digits m n).mapIdx fun i _ ↦ (f m) ^ i).sum := list_mul_sum (m.digits n) (f m) m + _ = m * ((f m ^ (d + 1) - 1) / (f m - 1)) := by + rw [list_geom _ (ne_of_gt (one_lt_of_not_bounded notbdd hm)), + (Nat.digits_len m n hm (not_eq_zero_of_lt hn)).symm] + _ ≤ m * ((f m ^ (d + 1))/(f m - 1)) := by + gcongr + · linarith only [one_lt_of_not_bounded notbdd hm] + · simp only [tsub_le_iff_right, le_add_iff_nonneg_right, zero_le_one] + _ = ↑m * f ↑m / (f ↑m - 1) * f ↑m ^ d := by ring + _ ≤ ↑m * f ↑m / (f ↑m - 1) * f ↑m ^ logb ↑m ↑n := by + gcongr + · exact le_of_lt (expr_pos hm notbdd) + · rw [← Real.rpow_natCast, Real.rpow_le_rpow_left_iff (one_lt_of_not_bounded notbdd hm)] + exact natLog_le_logb n m + apply le_of_pow_le_pow_left hk (mul_nonneg (rpow_nonneg + (le_of_lt (expr_pos hm notbdd)) (k : ℝ)⁻¹) (rpow_nonneg (apply_nonneg f ↑m) (logb m n))) + nth_rw 2 [← Real.rpow_natCast] + rw [mul_rpow (rpow_nonneg (le_of_lt (expr_pos hm notbdd)) (k : ℝ)⁻¹) + (rpow_nonneg (apply_nonneg f ↑m) (logb ↑m ↑n)), ← rpow_mul (le_of_lt (expr_pos hm notbdd)), + ← rpow_mul (apply_nonneg f ↑m), inv_mul_cancel₀ (mod_cast hk), rpow_one, mul_comm (logb ↑m ↑n)] + calc + (f n) ^ k = f ↑(n ^ k) := by simp only [Nat.cast_pow, map_pow] + _ ≤ (m * f m / (f m - 1)) * (f m) ^ (logb m ↑(n ^ k)) := h_ineq1 hm (Nat.one_lt_pow hk hn) + _ = (m * f m / (f m - 1)) * (f m) ^ (k * logb m n) := by + rw [Nat.cast_pow, Real.logb_pow] + exact_mod_cast zero_lt_of_lt hn + +include hm hn notbdd in +/-- Given two natural numbers `n, m` greater than 1 we have `f n ≤ f m ^ logb m n`. -/ +lemma mulRingNorm_le_mulRingNorm_pow_log : f n ≤ f m ^ logb m n := by + have : Tendsto (fun k : ℕ ↦ (m * f m / (f m - 1)) ^ (k : ℝ)⁻¹ * (f m) ^ (logb m n)) + atTop (𝓝 ((f m) ^ (logb m n))) := by + nth_rw 2 [← one_mul (f ↑m ^ logb ↑m ↑n)] + exact Tendsto.mul_const _ (tendsto_root_atTop_nhds_one (expr_pos hm notbdd)) + exact le_of_tendsto_of_tendsto (tendsto_const_nhds (x:= f ↑n)) this ((eventually_atTop.2 ⟨2, + fun b hb ↦ param_upperbound hm hn notbdd (not_eq_zero_of_lt hb)⟩)) + +include hm hn notbdd in +/-- Given m,n ≥ 2 and `f m = m ^ s`, `f n = n ^ t` for `s, t > 0`, we have `t ≤ s`. -/ +lemma le_of_mulRingNorm_eq {s t : ℝ} (hfm : f m = m ^ s) (hfn : f n = n ^ t) : t ≤ s := by + have hmn : f n ≤ f m ^ Real.logb m n := mulRingNorm_le_mulRingNorm_pow_log hm hn notbdd + rw [← Real.rpow_le_rpow_left_iff (x:=n) (mod_cast hn), ← hfn] + apply le_trans hmn + rw [hfm, ← Real.rpow_mul (Nat.cast_nonneg m), mul_comm, Real.rpow_mul (Nat.cast_nonneg m), + Real.rpow_logb (mod_cast zero_lt_of_lt hm) (mod_cast Nat.ne_of_lt' hm) + (mod_cast zero_lt_of_lt hn)] + +include hm hn notbdd in +private lemma symmetric_roles {s t : ℝ} (hfm : f m = m ^ s) (hfn : f n = n ^ t) : s = t := + le_antisymm (le_of_mulRingNorm_eq hn hm notbdd hfn hfm) + (le_of_mulRingNorm_eq hm hn notbdd hfm hfn) + +/-! ## Archimedean case: end goal -/ + +include notbdd in +/-- If `f` is not bounded and not trivial, then it is equivalent to the standard absolute value on +`ℚ`. -/ +theorem mulRingNorm_equiv_standard_of_unbounded : MulRingNorm.equiv f mulRingNorm_real := by + obtain ⟨m, hm⟩ := Classical.exists_not_of_not_forall notbdd + have oneltm : 1 < m := by + by_contra! + apply hm + replace this : m = 0 ∨ m = 1 := by omega + rcases this with (rfl | rfl) + all_goals simp only [CharP.cast_eq_zero, map_zero, zero_le_one, Nat.cast_one, map_one, le_refl] + rw [← equiv_on_nat_iff_equiv] + set s := Real.logb m (f m) with hs + use s⁻¹ + refine ⟨inv_pos.2 (Real.logb_pos (Nat.one_lt_cast.2 oneltm) + (one_lt_of_not_bounded notbdd oneltm)), ?_⟩ + intro n + by_cases h1 : n ≤ 1 + · by_cases h2 : n = 1 + · simp only [h2, Nat.cast_one, map_one, one_rpow, abs_one, cast_one] + · have : n = 0 := by omega + rw [this, hs] + simp only [CharP.cast_eq_zero, map_zero] + rw [Real.rpow_eq_zero le_rfl] + simp only [ne_eq, inv_eq_zero, logb_eq_zero, Nat.cast_eq_zero, Nat.cast_eq_one, map_eq_zero, + not_or] + push_neg + exact ⟨not_eq_zero_of_lt oneltm, Nat.ne_of_lt' oneltm, mod_cast (fun a ↦ a), + not_eq_zero_of_lt oneltm, ne_of_not_le hm, by linarith only [apply_nonneg f ↑m]⟩ + · simp only [mul_ring_norm_eq_abs, abs_cast, Rat.cast_natCast] + rw [Real.rpow_inv_eq (apply_nonneg f ↑n) (Nat.cast_nonneg n) + (Real.logb_ne_zero_of_pos_of_ne_one (one_lt_cast.mpr oneltm) (by linarith only [hm]) + (by linarith only [hm]))] + simp only [not_le] at h1 + have hfm : f m = m ^ s := by rw [Real.rpow_logb (mod_cast zero_lt_of_lt oneltm) + (mod_cast Nat.ne_of_lt' oneltm) (by linarith only [hm])] + have hfn : f n = n ^ (Real.logb n (f n)) := by + rw [Real.rpow_logb (mod_cast zero_lt_of_lt h1) (mod_cast Nat.ne_of_lt' h1) + (by apply map_pos_of_ne_zero; exact_mod_cast not_eq_zero_of_lt h1)] + rwa [← hs, symmetric_roles oneltm h1 notbdd hfm hfn] + end Archimedean +/-- Ostrowski's Theorem -/ +theorem mulRingNorm_equiv_standard_or_padic (f : MulRingNorm ℚ) (hf_nontriv : f ≠ 1) : + (MulRingNorm.equiv f mulRingNorm_real) ∨ + ∃! p, ∃ (hp : Fact (p.Prime)), MulRingNorm.equiv f (mulRingNorm_padic p) := by + by_cases bdd : ∀ n : ℕ, f n ≤ 1 + · right + exact mulRingNorm_equiv_padic_of_bounded hf_nontriv bdd + · left + exact mulRingNorm_equiv_standard_of_unbounded bdd + end Rat.MulRingNorm From 8a8d36d50901112a67ae1cfa45acd62a90bafa3b Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Mon, 21 Oct 2024 09:28:58 +0000 Subject: [PATCH 226/505] feat(GroupTheory/Index): index of product of subgroups (#17878) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Add two lemmas about the index of a product (pairwise or indexed) of subgroups. ```lean @[to_additive] lemma index_prod (H : Subgroup G) (K : Subgroup G') : (H.prod K).index = H.index * K.index := by ``` ```lean @[to_additive] lemma index_pi {ι : Type*} [Fintype ι] (H : ι → Subgroup G) : (Subgroup.pi Set.univ H).index = ∏ i, (H i).index := by ``` From AperiodicMonotilesLean. --- Mathlib/GroupTheory/Index.lean | 15 +++++++++++++++ 1 file changed, 15 insertions(+) diff --git a/Mathlib/GroupTheory/Index.lean b/Mathlib/GroupTheory/Index.lean index 977fdac89d177..237170a95e0dc 100644 --- a/Mathlib/GroupTheory/Index.lean +++ b/Mathlib/GroupTheory/Index.lean @@ -449,6 +449,21 @@ lemma pow_mem_of_relindex_ne_zero_of_dvd (h : H.relindex K ≠ 0) {a : G} (ha : convert pow_mem_of_index_ne_zero_of_dvd h ⟨a, ha⟩ hn simp [pow_mem ha, mem_subgroupOf] +@[to_additive (attr := simp)] +lemma index_prod (H : Subgroup G) (K : Subgroup G') : (H.prod K).index = H.index * K.index := by + simp_rw [index, ← Nat.card_prod] + refine Nat.card_congr + ((Quotient.congrRight (fun x y ↦ ?_)).trans (Setoid.prodQuotientEquiv _ _).symm) + rw [QuotientGroup.leftRel_prod] + +@[to_additive (attr := simp)] +lemma index_pi {ι : Type*} [Fintype ι] (H : ι → Subgroup G) : + (Subgroup.pi Set.univ H).index = ∏ i, (H i).index := by + simp_rw [index, ← Nat.card_pi] + refine Nat.card_congr + ((Quotient.congrRight (fun x y ↦ ?_)).trans (Setoid.piQuotientEquiv _).symm) + rw [QuotientGroup.leftRel_pi] + @[simp] lemma index_toAddSubgroup : (Subgroup.toAddSubgroup H).index = H.index := rfl From 662a8bae724a1bc23eb7626d0a6ce80a26fc14d5 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Mon, 21 Oct 2024 09:29:00 +0000 Subject: [PATCH 227/505] chore(Data/NNReal): split into `Defs` and `Basic` (#17913) I saw a few weird imports in `Data.NNReal.Basic` so I decided to split off a `Defs` file. `Defs` conains the definition and the conditionally complete linear ordered archimedean commutative semifield(!) structure, and all easily definable lemmas downstream of that. `Basic` contains whatever was left over and is a good target for further cleanup and reorganization. --- Mathlib.lean | 1 + Mathlib/Analysis/Normed/Group/Seminorm.lean | 3 +- Mathlib/Data/ENNReal/Basic.lean | 4 +- Mathlib/Data/ENNReal/Operations.lean | 1 + Mathlib/Data/Int/WithZero.lean | 2 +- Mathlib/Data/NNReal/Basic.lean | 957 +---------------- Mathlib/Data/NNReal/Defs.lean | 999 ++++++++++++++++++ Mathlib/Data/NNReal/Star.lean | 2 +- Mathlib/MeasureTheory/Measure/AddContent.lean | 1 + Mathlib/RingTheory/Valuation/RankOne.lean | 2 +- 10 files changed, 1015 insertions(+), 957 deletions(-) create mode 100644 Mathlib/Data/NNReal/Defs.lean diff --git a/Mathlib.lean b/Mathlib.lean index 7c599ec0aa9d0..6da8bfa58c2fb 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2437,6 +2437,7 @@ import Mathlib.Data.NNRat.Floor import Mathlib.Data.NNRat.Lemmas import Mathlib.Data.NNRat.Order import Mathlib.Data.NNReal.Basic +import Mathlib.Data.NNReal.Defs import Mathlib.Data.NNReal.Star import Mathlib.Data.Nat.BinaryRec import Mathlib.Data.Nat.BitIndices diff --git a/Mathlib/Analysis/Normed/Group/Seminorm.lean b/Mathlib/Analysis/Normed/Group/Seminorm.lean index a8d9220ba89cc..73c549f43211b 100644 --- a/Mathlib/Analysis/Normed/Group/Seminorm.lean +++ b/Mathlib/Analysis/Normed/Group/Seminorm.lean @@ -3,7 +3,8 @@ Copyright (c) 2022 María Inés de Frutos-Fernández, Yaël Dillies. All rights Released under Apache 2.0 license as described in the file LICENSE. Authors: María Inés de Frutos-Fernández, Yaël Dillies -/ -import Mathlib.Data.NNReal.Basic +import Mathlib.Data.NNReal.Defs +import Mathlib.Order.ConditionallyCompleteLattice.Group import Mathlib.Tactic.GCongr.CoreAttrs /-! diff --git a/Mathlib/Data/ENNReal/Basic.lean b/Mathlib/Data/ENNReal/Basic.lean index cd2388d50e256..0f906e0ef21a5 100644 --- a/Mathlib/Data/ENNReal/Basic.lean +++ b/Mathlib/Data/ENNReal/Basic.lean @@ -3,9 +3,11 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Yury Kudryashov -/ +import Mathlib.Algebra.Group.Indicator import Mathlib.Algebra.Order.Ring.WithTop import Mathlib.Algebra.Order.Sub.WithTop -import Mathlib.Data.NNReal.Basic +import Mathlib.Data.Finset.Lattice +import Mathlib.Data.NNReal.Defs import Mathlib.Order.Interval.Set.WithBotTop /-! diff --git a/Mathlib/Data/ENNReal/Operations.lean b/Mathlib/Data/ENNReal/Operations.lean index 8318373b6a1bf..6931936b8b2fe 100644 --- a/Mathlib/Data/ENNReal/Operations.lean +++ b/Mathlib/Data/ENNReal/Operations.lean @@ -6,6 +6,7 @@ Authors: Johannes Hölzl, Yury Kudryashov import Mathlib.Algebra.BigOperators.WithTop import Mathlib.Algebra.GroupWithZero.Divisibility import Mathlib.Data.ENNReal.Basic +import Mathlib.Data.NNReal.Basic /-! # Properties of addition, multiplication and subtraction on extended non-negative real numbers diff --git a/Mathlib/Data/Int/WithZero.lean b/Mathlib/Data/Int/WithZero.lean index ec6f9a4c79ab0..782b0c9336573 100644 --- a/Mathlib/Data/Int/WithZero.lean +++ b/Mathlib/Data/Int/WithZero.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 María Inés de Frutos-Fernández, Filippo A. E. Nuccio. All Released under Apache 2.0 license as described in the file LICENSE. Authors: María Inés de Frutos-Fernández, Filippo A. E. Nuccio -/ -import Mathlib.Data.NNReal.Basic +import Mathlib.Data.NNReal.Defs /-! # WithZero diff --git a/Mathlib/Data/NNReal/Basic.lean b/Mathlib/Data/NNReal/Basic.lean index 0db61f23290b1..be8f76fd33712 100644 --- a/Mathlib/Data/NNReal/Basic.lean +++ b/Mathlib/Data/NNReal/Basic.lean @@ -7,46 +7,19 @@ import Mathlib.Algebra.BigOperators.Expect import Mathlib.Algebra.Order.BigOperators.Ring.Finset import Mathlib.Algebra.Order.Field.Canonical.Basic import Mathlib.Algebra.Order.Nonneg.Floor -import Mathlib.Algebra.Ring.Regular import Mathlib.Data.Real.Pointwise +import Mathlib.Data.NNReal.Defs import Mathlib.Order.ConditionallyCompleteLattice.Group /-! -# Nonnegative real numbers +# Basic results on nonnegative real numbers -In this file we define `NNReal` (notation: `ℝ≥0`) to be the type of non-negative real numbers, -a.k.a. the interval `[0, ∞)`. We also define the following operations and structures on `ℝ≥0`: - -* the order on `ℝ≥0` is the restriction of the order on `ℝ`; these relations define a conditionally - complete linear order with a bottom element, `ConditionallyCompleteLinearOrderBot`; - -* `a + b` and `a * b` are the restrictions of addition and multiplication of real numbers to `ℝ≥0`; - these operations together with `0 = ⟨0, _⟩` and `1 = ⟨1, _⟩` turn `ℝ≥0` into a conditionally - complete linear ordered archimedean commutative semifield; we have no typeclass for this in - `mathlib` yet, so we define the following instances instead: - - - `LinearOrderedSemiring ℝ≥0`; - - `OrderedCommSemiring ℝ≥0`; - - `CanonicallyOrderedCommSemiring ℝ≥0`; - - `LinearOrderedCommGroupWithZero ℝ≥0`; - - `CanonicallyLinearOrderedAddCommMonoid ℝ≥0`; - - `Archimedean ℝ≥0`; - - `ConditionallyCompleteLinearOrderBot ℝ≥0`. - - These instances are derived from corresponding instances about the type `{x : α // 0 ≤ x}` in an - appropriate ordered field/ring/group/monoid `α`, see `Mathlib.Algebra.Order.Nonneg.OrderedRing`. - -* `Real.toNNReal x` is defined as `⟨max x 0, _⟩`, i.e. `↑(Real.toNNReal x) = x` when `0 ≤ x` and - `↑(Real.toNNReal x) = 0` otherwise. - -We also define an instance `CanLift ℝ ℝ≥0`. This instance can be used by the `lift` tactic to -replace `x : ℝ` and `hx : 0 ≤ x` in the proof context with `x : ℝ≥0` while replacing all occurrences -of `x` with `↑x`. This tactic also works for a function `f : α → ℝ` with a hypothesis -`hf : ∀ x, 0 ≤ f x`. +This file contains all results on `NNReal` that do not directly follow from its basic structure. +As a consequence, it is a bit of a random collection of results, and is a good target for cleanup. ## Notations -This file defines `ℝ≥0` as a localized notation for `NNReal`. +This file uses `ℝ≥0` as a localized notation for `NNReal`. -/ assert_not_exists Star @@ -54,212 +27,15 @@ assert_not_exists Star open Function open scoped BigOperators --- to ensure these instances are computable -/-- Nonnegative real numbers. -/ -def NNReal := { r : ℝ // 0 ≤ r } deriving - Zero, One, Semiring, StrictOrderedSemiring, CommMonoidWithZero, CommSemiring, - SemilatticeInf, SemilatticeSup, DistribLattice, OrderedCommSemiring, - CanonicallyOrderedCommSemiring, Inhabited - namespace NNReal -scoped notation "ℝ≥0" => NNReal - noncomputable instance : FloorSemiring ℝ≥0 := Nonneg.floorSemiring -instance instDenselyOrdered : DenselyOrdered ℝ≥0 := Nonneg.instDenselyOrdered -instance : OrderBot ℝ≥0 := inferInstance -instance instArchimedean : Archimedean ℝ≥0 := Nonneg.instArchimedean -instance instMulArchimedean : MulArchimedean ℝ≥0 := Nonneg.instMulArchimedean -noncomputable instance : Sub ℝ≥0 := Nonneg.sub -noncomputable instance : OrderedSub ℝ≥0 := Nonneg.orderedSub - -noncomputable instance : CanonicallyLinearOrderedSemifield ℝ≥0 := - Nonneg.canonicallyLinearOrderedSemifield - -/-- Coercion `ℝ≥0 → ℝ`. -/ -@[coe] def toReal : ℝ≥0 → ℝ := Subtype.val - -instance : Coe ℝ≥0 ℝ := ⟨toReal⟩ - --- Simp lemma to put back `n.val` into the normal form given by the coercion. -@[simp] -theorem val_eq_coe (n : ℝ≥0) : n.val = n := - rfl - -instance canLift : CanLift ℝ ℝ≥0 toReal fun r => 0 ≤ r := - Subtype.canLift _ - -@[ext] protected theorem eq {n m : ℝ≥0} : (n : ℝ) = (m : ℝ) → n = m := - Subtype.eq - -theorem ne_iff {x y : ℝ≥0} : (x : ℝ) ≠ (y : ℝ) ↔ x ≠ y := - not_congr <| NNReal.eq_iff.symm - -protected theorem «forall» {p : ℝ≥0 → Prop} : - (∀ x : ℝ≥0, p x) ↔ ∀ (x : ℝ) (hx : 0 ≤ x), p ⟨x, hx⟩ := - Subtype.forall - -protected theorem «exists» {p : ℝ≥0 → Prop} : - (∃ x : ℝ≥0, p x) ↔ ∃ (x : ℝ) (hx : 0 ≤ x), p ⟨x, hx⟩ := - Subtype.exists - -/-- Reinterpret a real number `r` as a non-negative real number. Returns `0` if `r < 0`. -/ -noncomputable def _root_.Real.toNNReal (r : ℝ) : ℝ≥0 := - ⟨max r 0, le_max_right _ _⟩ - -theorem _root_.Real.coe_toNNReal (r : ℝ) (hr : 0 ≤ r) : (Real.toNNReal r : ℝ) = r := - max_eq_left hr - -theorem _root_.Real.toNNReal_of_nonneg {r : ℝ} (hr : 0 ≤ r) : r.toNNReal = ⟨r, hr⟩ := by - simp_rw [Real.toNNReal, max_eq_left hr] - -theorem _root_.Real.le_coe_toNNReal (r : ℝ) : r ≤ Real.toNNReal r := - le_max_left r 0 - -@[bound] theorem coe_nonneg (r : ℝ≥0) : (0 : ℝ) ≤ r := r.2 - -@[simp, norm_cast] theorem coe_mk (a : ℝ) (ha) : toReal ⟨a, ha⟩ = a := rfl - -example : Zero ℝ≥0 := by infer_instance - -example : One ℝ≥0 := by infer_instance - -example : Add ℝ≥0 := by infer_instance - -noncomputable example : Sub ℝ≥0 := by infer_instance - -example : Mul ℝ≥0 := by infer_instance - -noncomputable example : Inv ℝ≥0 := by infer_instance - -noncomputable example : Div ℝ≥0 := by infer_instance - -example : LE ℝ≥0 := by infer_instance - -example : Bot ℝ≥0 := by infer_instance - -example : Inhabited ℝ≥0 := by infer_instance - -example : Nontrivial ℝ≥0 := by infer_instance - -protected theorem coe_injective : Injective ((↑) : ℝ≥0 → ℝ) := Subtype.coe_injective - -@[simp, norm_cast] lemma coe_inj {r₁ r₂ : ℝ≥0} : (r₁ : ℝ) = r₂ ↔ r₁ = r₂ := - NNReal.coe_injective.eq_iff - -@[deprecated (since := "2024-02-03")] protected alias coe_eq := coe_inj - -@[simp, norm_cast] lemma coe_zero : ((0 : ℝ≥0) : ℝ) = 0 := rfl - -@[simp, norm_cast] lemma coe_one : ((1 : ℝ≥0) : ℝ) = 1 := rfl - -@[simp] lemma mk_zero : (⟨0, le_rfl⟩ : ℝ≥0) = 0 := rfl -@[simp] lemma mk_one : (⟨1, zero_le_one⟩ : ℝ≥0) = 1 := rfl - -@[simp, norm_cast] -protected theorem coe_add (r₁ r₂ : ℝ≥0) : ((r₁ + r₂ : ℝ≥0) : ℝ) = r₁ + r₂ := - rfl - -@[simp, norm_cast] -protected theorem coe_mul (r₁ r₂ : ℝ≥0) : ((r₁ * r₂ : ℝ≥0) : ℝ) = r₁ * r₂ := - rfl - -@[simp, norm_cast] -protected theorem coe_inv (r : ℝ≥0) : ((r⁻¹ : ℝ≥0) : ℝ) = (r : ℝ)⁻¹ := - rfl - -@[simp, norm_cast] -protected theorem coe_div (r₁ r₂ : ℝ≥0) : ((r₁ / r₂ : ℝ≥0) : ℝ) = (r₁ : ℝ) / r₂ := - rfl - -protected theorem coe_two : ((2 : ℝ≥0) : ℝ) = 2 := rfl - -@[simp, norm_cast] -protected theorem coe_sub {r₁ r₂ : ℝ≥0} (h : r₂ ≤ r₁) : ((r₁ - r₂ : ℝ≥0) : ℝ) = ↑r₁ - ↑r₂ := - max_eq_left <| le_sub_comm.2 <| by simp [show (r₂ : ℝ) ≤ r₁ from h] - -variable {r r₁ r₂ : ℝ≥0} {x y : ℝ} - -@[simp, norm_cast] lemma coe_eq_zero : (r : ℝ) = 0 ↔ r = 0 := by rw [← coe_zero, coe_inj] - -@[simp, norm_cast] lemma coe_eq_one : (r : ℝ) = 1 ↔ r = 1 := by rw [← coe_one, coe_inj] - -@[norm_cast] lemma coe_ne_zero : (r : ℝ) ≠ 0 ↔ r ≠ 0 := coe_eq_zero.not - -@[norm_cast] lemma coe_ne_one : (r : ℝ) ≠ 1 ↔ r ≠ 1 := coe_eq_one.not - -example : CommSemiring ℝ≥0 := by infer_instance - -/-- Coercion `ℝ≥0 → ℝ` as a `RingHom`. - -Porting note (#11215): TODO: what if we define `Coe ℝ≥0 ℝ` using this function? -/ -def toRealHom : ℝ≥0 →+* ℝ where - toFun := (↑) - map_one' := NNReal.coe_one - map_mul' := NNReal.coe_mul - map_zero' := NNReal.coe_zero - map_add' := NNReal.coe_add - -@[simp] theorem coe_toRealHom : ⇑toRealHom = toReal := rfl - -section Actions - -/-- A `MulAction` over `ℝ` restricts to a `MulAction` over `ℝ≥0`. -/ -instance {M : Type*} [MulAction ℝ M] : MulAction ℝ≥0 M := - MulAction.compHom M toRealHom.toMonoidHom - -theorem smul_def {M : Type*} [MulAction ℝ M] (c : ℝ≥0) (x : M) : c • x = (c : ℝ) • x := - rfl - -instance {M N : Type*} [MulAction ℝ M] [MulAction ℝ N] [SMul M N] [IsScalarTower ℝ M N] : - IsScalarTower ℝ≥0 M N where smul_assoc r := (smul_assoc (r : ℝ) : _) - -instance smulCommClass_left {M N : Type*} [MulAction ℝ N] [SMul M N] [SMulCommClass ℝ M N] : - SMulCommClass ℝ≥0 M N where smul_comm r := (smul_comm (r : ℝ) : _) - -instance smulCommClass_right {M N : Type*} [MulAction ℝ N] [SMul M N] [SMulCommClass M ℝ N] : - SMulCommClass M ℝ≥0 N where smul_comm m r := (smul_comm m (r : ℝ) : _) - -/-- A `DistribMulAction` over `ℝ` restricts to a `DistribMulAction` over `ℝ≥0`. -/ -instance {M : Type*} [AddMonoid M] [DistribMulAction ℝ M] : DistribMulAction ℝ≥0 M := - DistribMulAction.compHom M toRealHom.toMonoidHom - -/-- A `Module` over `ℝ` restricts to a `Module` over `ℝ≥0`. -/ -instance {M : Type*} [AddCommMonoid M] [Module ℝ M] : Module ℝ≥0 M := - Module.compHom M toRealHom - --- Porting note (#11215): TODO: after this line, `↑` uses `Algebra.cast` instead of `toReal` -/-- An `Algebra` over `ℝ` restricts to an `Algebra` over `ℝ≥0`. -/ -instance {A : Type*} [Semiring A] [Algebra ℝ A] : Algebra ℝ≥0 A where - smul := (· • ·) - commutes' r x := by simp [Algebra.commutes] - smul_def' r x := by simp [← Algebra.smul_def (r : ℝ) x, smul_def] - toRingHom := (algebraMap ℝ A).comp (toRealHom : ℝ≥0 →+* ℝ) - --- verify that the above produces instances we might care about -example : Algebra ℝ≥0 ℝ := by infer_instance - -example : DistribMulAction ℝ≥0ˣ ℝ := by infer_instance - -end Actions - -example : MonoidWithZero ℝ≥0 := by infer_instance - -example : CommMonoidWithZero ℝ≥0 := by infer_instance - -noncomputable example : CommGroupWithZero ℝ≥0 := by infer_instance @[simp, norm_cast] theorem coe_indicator {α} (s : Set α) (f : α → ℝ≥0) (a : α) : ((s.indicator f a : ℝ≥0) : ℝ) = s.indicator (fun x => ↑(f x)) a := (toRealHom : ℝ≥0 →+ ℝ).map_indicator _ _ _ -@[simp, norm_cast] -theorem coe_pow (r : ℝ≥0) (n : ℕ) : ((r ^ n : ℝ≥0) : ℝ) = (r : ℝ) ^ n := rfl - -@[simp, norm_cast] -theorem coe_zpow (r : ℝ≥0) (n : ℤ) : ((r ^ n : ℝ≥0) : ℝ) = (r : ℝ) ^ n := rfl - @[norm_cast] theorem coe_list_sum (l : List ℝ≥0) : ((l.sum : ℝ≥0) : ℝ) = (l.map (↑)).sum := map_list_sum toRealHom l @@ -300,226 +76,11 @@ theorem _root_.Real.toNNReal_prod_of_nonneg (hf : ∀ a, a ∈ s → 0 ≤ f a) rw [← coe_inj, NNReal.coe_prod, Real.coe_toNNReal _ (Finset.prod_nonneg hf)] exact Finset.prod_congr rfl fun x hxs => by rw [Real.coe_toNNReal _ (hf x hxs)] -@[simp, norm_cast] lemma coe_nsmul (r : ℝ≥0) (n : ℕ) : ↑(n • r) = n • (r : ℝ) := rfl -@[simp, norm_cast] lemma coe_nnqsmul (q : ℚ≥0) (x : ℝ≥0) : ↑(q • x) = (q • x : ℝ) := rfl - -@[simp, norm_cast] -protected theorem coe_natCast (n : ℕ) : (↑(↑n : ℝ≥0) : ℝ) = n := - map_natCast toRealHom n - -@[deprecated (since := "2024-04-17")] -alias coe_nat_cast := NNReal.coe_natCast - --- See note [no_index around OfNat.ofNat] -@[simp, norm_cast] -protected theorem coe_ofNat (n : ℕ) [n.AtLeastTwo] : - (no_index (OfNat.ofNat n : ℝ≥0) : ℝ) = OfNat.ofNat n := - rfl - -@[simp, norm_cast] -protected theorem coe_ofScientific (m : ℕ) (s : Bool) (e : ℕ) : - ↑(OfScientific.ofScientific m s e : ℝ≥0) = (OfScientific.ofScientific m s e : ℝ) := - rfl - -@[simp, norm_cast] -lemma algebraMap_eq_coe : (algebraMap ℝ≥0 ℝ : ℝ≥0 → ℝ) = (↑) := rfl - -noncomputable example : LinearOrder ℝ≥0 := by infer_instance - -@[simp, norm_cast] lemma coe_le_coe : (r₁ : ℝ) ≤ r₂ ↔ r₁ ≤ r₂ := Iff.rfl - -@[simp, norm_cast] lemma coe_lt_coe : (r₁ : ℝ) < r₂ ↔ r₁ < r₂ := Iff.rfl - -@[bound] private alias ⟨_, Bound.coe_lt_coe_of_lt⟩ := coe_lt_coe - -@[simp, norm_cast] lemma coe_pos : (0 : ℝ) < r ↔ 0 < r := Iff.rfl - -@[bound] private alias ⟨_, Bound.coe_pos_of_pos⟩ := coe_pos - -@[simp, norm_cast] lemma one_le_coe : 1 ≤ (r : ℝ) ↔ 1 ≤ r := by rw [← coe_le_coe, coe_one] -@[simp, norm_cast] lemma one_lt_coe : 1 < (r : ℝ) ↔ 1 < r := by rw [← coe_lt_coe, coe_one] -@[simp, norm_cast] lemma coe_le_one : (r : ℝ) ≤ 1 ↔ r ≤ 1 := by rw [← coe_le_coe, coe_one] -@[simp, norm_cast] lemma coe_lt_one : (r : ℝ) < 1 ↔ r < 1 := by rw [← coe_lt_coe, coe_one] - -@[mono] lemma coe_mono : Monotone ((↑) : ℝ≥0 → ℝ) := fun _ _ => NNReal.coe_le_coe.2 - -/-- Alias for the use of `gcongr` -/ -@[gcongr] alias ⟨_, GCongr.toReal_le_toReal⟩ := coe_le_coe - -protected theorem _root_.Real.toNNReal_mono : Monotone Real.toNNReal := fun _ _ h => - max_le_max h (le_refl 0) - -@[simp] -theorem _root_.Real.toNNReal_coe {r : ℝ≥0} : Real.toNNReal r = r := - NNReal.eq <| max_eq_left r.2 - -@[simp] -theorem mk_natCast (n : ℕ) : @Eq ℝ≥0 (⟨(n : ℝ), n.cast_nonneg⟩ : ℝ≥0) n := - NNReal.eq (NNReal.coe_natCast n).symm - -@[deprecated (since := "2024-04-05")] alias mk_coe_nat := mk_natCast - --- Porting note: place this in the `Real` namespace -@[simp] -theorem toNNReal_coe_nat (n : ℕ) : Real.toNNReal n = n := - NNReal.eq <| by simp [Real.coe_toNNReal] - --- See note [no_index around OfNat.ofNat] -@[simp] -theorem _root_.Real.toNNReal_ofNat (n : ℕ) [n.AtLeastTwo] : - Real.toNNReal (no_index (OfNat.ofNat n)) = OfNat.ofNat n := - toNNReal_coe_nat n - -/-- `Real.toNNReal` and `NNReal.toReal : ℝ≥0 → ℝ` form a Galois insertion. -/ -noncomputable def gi : GaloisInsertion Real.toNNReal (↑) := - GaloisInsertion.monotoneIntro NNReal.coe_mono Real.toNNReal_mono Real.le_coe_toNNReal fun _ => - Real.toNNReal_coe - --- note that anything involving the (decidability of the) linear order, --- will be noncomputable, everything else should not be. -example : OrderBot ℝ≥0 := by infer_instance - -example : PartialOrder ℝ≥0 := by infer_instance - -noncomputable example : CanonicallyLinearOrderedAddCommMonoid ℝ≥0 := by infer_instance - -noncomputable example : LinearOrderedAddCommMonoid ℝ≥0 := by infer_instance - -example : DistribLattice ℝ≥0 := by infer_instance - -example : SemilatticeInf ℝ≥0 := by infer_instance - -example : SemilatticeSup ℝ≥0 := by infer_instance - -noncomputable example : LinearOrderedSemiring ℝ≥0 := by infer_instance - -example : OrderedCommSemiring ℝ≥0 := by infer_instance - -noncomputable example : LinearOrderedCommMonoid ℝ≥0 := by infer_instance - -noncomputable example : LinearOrderedCommMonoidWithZero ℝ≥0 := by infer_instance - -noncomputable example : LinearOrderedCommGroupWithZero ℝ≥0 := by infer_instance - -example : CanonicallyOrderedCommSemiring ℝ≥0 := by infer_instance - -example : DenselyOrdered ℝ≥0 := by infer_instance - -example : NoMaxOrder ℝ≥0 := by infer_instance - -instance instPosSMulStrictMono {α} [Preorder α] [MulAction ℝ α] [PosSMulStrictMono ℝ α] : - PosSMulStrictMono ℝ≥0 α where - elim _r hr _a₁ _a₂ ha := (smul_lt_smul_of_pos_left ha (coe_pos.2 hr):) - -instance instSMulPosStrictMono {α} [Zero α] [Preorder α] [MulAction ℝ α] [SMulPosStrictMono ℝ α] : - SMulPosStrictMono ℝ≥0 α where - elim _a ha _r₁ _r₂ hr := (smul_lt_smul_of_pos_right (coe_lt_coe.2 hr) ha :) - -/-- If `a` is a nonnegative real number, then the closed interval `[0, a]` in `ℝ` is order -isomorphic to the interval `Set.Iic a`. -/ --- Porting note (#11215): TODO: restore once `simps` supports `ℝ≥0` @[simps!? apply_coe_coe] -def orderIsoIccZeroCoe (a : ℝ≥0) : Set.Icc (0 : ℝ) a ≃o Set.Iic a where - toEquiv := Equiv.Set.sep (Set.Ici 0) fun x : ℝ => x ≤ a - map_rel_iff' := Iff.rfl - -@[simp] -theorem orderIsoIccZeroCoe_apply_coe_coe (a : ℝ≥0) (b : Set.Icc (0 : ℝ) a) : - (orderIsoIccZeroCoe a b : ℝ) = b := - rfl - -@[simp] -theorem orderIsoIccZeroCoe_symm_apply_coe (a : ℝ≥0) (b : Set.Iic a) : - ((orderIsoIccZeroCoe a).symm b : ℝ) = b := - rfl - --- note we need the `@` to make the `Membership.mem` have a sensible type -theorem coe_image {s : Set ℝ≥0} : - (↑) '' s = { x : ℝ | ∃ h : 0 ≤ x, @Membership.mem ℝ≥0 _ _ s ⟨x, h⟩ } := - Subtype.coe_image - -theorem bddAbove_coe {s : Set ℝ≥0} : BddAbove (((↑) : ℝ≥0 → ℝ) '' s) ↔ BddAbove s := - Iff.intro - (fun ⟨b, hb⟩ => - ⟨Real.toNNReal b, fun ⟨y, _⟩ hys => - show y ≤ max b 0 from le_max_of_le_left <| hb <| Set.mem_image_of_mem _ hys⟩) - fun ⟨b, hb⟩ => ⟨b, fun _ ⟨_, hx, eq⟩ => eq ▸ hb hx⟩ - -theorem bddBelow_coe (s : Set ℝ≥0) : BddBelow (((↑) : ℝ≥0 → ℝ) '' s) := - ⟨0, fun _ ⟨q, _, eq⟩ => eq ▸ q.2⟩ - -noncomputable instance : ConditionallyCompleteLinearOrderBot ℝ≥0 := - Nonneg.conditionallyCompleteLinearOrderBot 0 - -@[norm_cast] -theorem coe_sSup (s : Set ℝ≥0) : (↑(sSup s) : ℝ) = sSup (((↑) : ℝ≥0 → ℝ) '' s) := by - rcases Set.eq_empty_or_nonempty s with rfl|hs - · simp - by_cases H : BddAbove s - · have A : sSup (Subtype.val '' s) ∈ Set.Ici 0 := by - apply Real.sSup_nonneg - rintro - ⟨y, -, rfl⟩ - exact y.2 - exact (@subset_sSup_of_within ℝ (Set.Ici (0 : ℝ)) _ _ (_) s hs H A).symm - · simp only [csSup_of_not_bddAbove H, csSup_empty, bot_eq_zero', NNReal.coe_zero] - apply (Real.sSup_of_not_bddAbove ?_).symm - contrapose! H - exact bddAbove_coe.1 H - -@[simp, norm_cast] -- Porting note: add `simp` -theorem coe_iSup {ι : Sort*} (s : ι → ℝ≥0) : (↑(⨆ i, s i) : ℝ) = ⨆ i, ↑(s i) := by - rw [iSup, iSup, coe_sSup, ← Set.range_comp]; rfl - -@[norm_cast] -theorem coe_sInf (s : Set ℝ≥0) : (↑(sInf s) : ℝ) = sInf (((↑) : ℝ≥0 → ℝ) '' s) := by - rcases Set.eq_empty_or_nonempty s with rfl|hs - · simp only [Set.image_empty, Real.sInf_empty, coe_eq_zero] - exact @subset_sInf_emptyset ℝ (Set.Ici (0 : ℝ)) _ _ (_) - have A : sInf (Subtype.val '' s) ∈ Set.Ici 0 := by - apply Real.sInf_nonneg - rintro - ⟨y, -, rfl⟩ - exact y.2 - exact (@subset_sInf_of_within ℝ (Set.Ici (0 : ℝ)) _ _ (_) s hs (OrderBot.bddBelow s) A).symm - -@[simp] -theorem sInf_empty : sInf (∅ : Set ℝ≥0) = 0 := by - rw [← coe_eq_zero, coe_sInf, Set.image_empty, Real.sInf_empty] - -@[norm_cast] -theorem coe_iInf {ι : Sort*} (s : ι → ℝ≥0) : (↑(⨅ i, s i) : ℝ) = ⨅ i, ↑(s i) := by - rw [iInf, iInf, coe_sInf, ← Set.range_comp]; rfl - theorem le_iInf_add_iInf {ι ι' : Sort*} [Nonempty ι] [Nonempty ι'] {f : ι → ℝ≥0} {g : ι' → ℝ≥0} {a : ℝ≥0} (h : ∀ i j, a ≤ f i + g j) : a ≤ (⨅ i, f i) + ⨅ j, g j := by rw [← NNReal.coe_le_coe, NNReal.coe_add, coe_iInf, coe_iInf] exact le_ciInf_add_ciInf h --- Short-circuit instance search -instance instCovariantClassAddLE : CovariantClass ℝ≥0 ℝ≥0 (· + ·) (· ≤ ·) := inferInstance -instance instContravariantClassAddLT : ContravariantClass ℝ≥0 ℝ≥0 (· + ·) (· < ·) := inferInstance -instance instCovariantClassMulLE : CovariantClass ℝ≥0 ℝ≥0 (· * ·) (· ≤ ·) := inferInstance - -@[deprecated le_of_forall_pos_le_add (since := "2024-10-17")] -protected theorem le_of_forall_pos_le_add {a b : ℝ≥0} (h : ∀ ε, 0 < ε → a ≤ b + ε) : a ≤ b := - le_of_forall_pos_le_add h - -theorem lt_iff_exists_rat_btwn (a b : ℝ≥0) : - a < b ↔ ∃ q : ℚ, 0 ≤ q ∧ a < Real.toNNReal q ∧ Real.toNNReal q < b := - Iff.intro - (fun h : (↑a : ℝ) < (↑b : ℝ) => - let ⟨q, haq, hqb⟩ := exists_rat_btwn h - have : 0 ≤ (q : ℝ) := le_trans a.2 <| le_of_lt haq - ⟨q, Rat.cast_nonneg.1 this, by - simp [Real.coe_toNNReal _ this, NNReal.coe_lt_coe.symm, haq, hqb]⟩) - fun ⟨_, _, haq, hqb⟩ => lt_trans haq hqb - -theorem bot_eq_zero : (⊥ : ℝ≥0) = 0 := rfl - -theorem mul_sup (a b c : ℝ≥0) : a * (b ⊔ c) = a * b ⊔ a * c := - mul_max_of_nonneg _ _ <| zero_le a - -theorem sup_mul (a b c : ℝ≥0) : (a ⊔ b) * c = a * c ⊔ b * c := - max_mul_of_nonneg _ _ <| zero_le c - theorem mul_finset_sup {α} (r : ℝ≥0) (s : Finset α) (f : α → ℝ≥0) : r * s.sup f = s.sup fun a => r * f a := Finset.comp_sup_eq_sup_comp _ (NNReal.mul_sup r) (mul_zero r) @@ -531,248 +92,8 @@ theorem finset_sup_mul {α} (s : Finset α) (f : α → ℝ≥0) (r : ℝ≥0) : theorem finset_sup_div {α} {f : α → ℝ≥0} {s : Finset α} (r : ℝ≥0) : s.sup f / r = s.sup fun a => f a / r := by simp only [div_eq_inv_mul, mul_finset_sup] -@[simp, norm_cast] -theorem coe_max (x y : ℝ≥0) : ((max x y : ℝ≥0) : ℝ) = max (x : ℝ) (y : ℝ) := - NNReal.coe_mono.map_max - -@[simp, norm_cast] -theorem coe_min (x y : ℝ≥0) : ((min x y : ℝ≥0) : ℝ) = min (x : ℝ) (y : ℝ) := - NNReal.coe_mono.map_min - -@[simp] -theorem zero_le_coe {q : ℝ≥0} : 0 ≤ (q : ℝ) := - q.2 - -instance instOrderedSMul {M : Type*} [OrderedAddCommMonoid M] [Module ℝ M] [OrderedSMul ℝ M] : - OrderedSMul ℝ≥0 M where - smul_lt_smul_of_pos hab hc := (smul_lt_smul_of_pos_left hab (NNReal.coe_pos.2 hc) : _) - lt_of_smul_lt_smul_of_pos {_ _ c} hab _ := - lt_of_smul_lt_smul_of_nonneg_left (by exact hab) (NNReal.coe_nonneg c) - -end NNReal - -open NNReal - -namespace Real - -section ToNNReal - -@[simp] -theorem coe_toNNReal' (r : ℝ) : (Real.toNNReal r : ℝ) = max r 0 := - rfl - -@[simp] -theorem toNNReal_zero : Real.toNNReal 0 = 0 := NNReal.eq <| coe_toNNReal _ le_rfl - -@[simp] -theorem toNNReal_one : Real.toNNReal 1 = 1 := NNReal.eq <| coe_toNNReal _ zero_le_one - -@[simp] -theorem toNNReal_pos {r : ℝ} : 0 < Real.toNNReal r ↔ 0 < r := by - simp [← NNReal.coe_lt_coe, lt_irrefl] - -@[simp] -theorem toNNReal_eq_zero {r : ℝ} : Real.toNNReal r = 0 ↔ r ≤ 0 := by - simpa [-toNNReal_pos] using not_iff_not.2 (@toNNReal_pos r) - -theorem toNNReal_of_nonpos {r : ℝ} : r ≤ 0 → Real.toNNReal r = 0 := - toNNReal_eq_zero.2 - -lemma toNNReal_eq_iff_eq_coe {r : ℝ} {p : ℝ≥0} (hp : p ≠ 0) : r.toNNReal = p ↔ r = p := - ⟨fun h ↦ h ▸ (coe_toNNReal _ <| not_lt.1 fun hlt ↦ hp <| h ▸ toNNReal_of_nonpos hlt.le).symm, - fun h ↦ h.symm ▸ toNNReal_coe⟩ - -@[simp] -lemma toNNReal_eq_one {r : ℝ} : r.toNNReal = 1 ↔ r = 1 := toNNReal_eq_iff_eq_coe one_ne_zero - -@[simp] -lemma toNNReal_eq_natCast {r : ℝ} {n : ℕ} (hn : n ≠ 0) : r.toNNReal = n ↔ r = n := - mod_cast toNNReal_eq_iff_eq_coe <| Nat.cast_ne_zero.2 hn - -@[deprecated (since := "2024-04-17")] -alias toNNReal_eq_nat_cast := toNNReal_eq_natCast - -@[simp] -lemma toNNReal_eq_ofNat {r : ℝ} {n : ℕ} [n.AtLeastTwo] : - r.toNNReal = no_index (OfNat.ofNat n) ↔ r = OfNat.ofNat n := - toNNReal_eq_natCast (NeZero.ne n) - -@[simp] -theorem toNNReal_le_toNNReal_iff {r p : ℝ} (hp : 0 ≤ p) : - toNNReal r ≤ toNNReal p ↔ r ≤ p := by simp [← NNReal.coe_le_coe, hp] - -@[simp] -lemma toNNReal_le_one {r : ℝ} : r.toNNReal ≤ 1 ↔ r ≤ 1 := by - simpa using toNNReal_le_toNNReal_iff zero_le_one - -@[simp] -lemma one_lt_toNNReal {r : ℝ} : 1 < r.toNNReal ↔ 1 < r := by - simpa only [not_le] using toNNReal_le_one.not - -@[simp] -lemma toNNReal_le_natCast {r : ℝ} {n : ℕ} : r.toNNReal ≤ n ↔ r ≤ n := by - simpa using toNNReal_le_toNNReal_iff n.cast_nonneg - -@[deprecated (since := "2024-04-17")] -alias toNNReal_le_nat_cast := toNNReal_le_natCast - -@[simp] -lemma natCast_lt_toNNReal {r : ℝ} {n : ℕ} : n < r.toNNReal ↔ n < r := by - simpa only [not_le] using toNNReal_le_natCast.not - -@[deprecated (since := "2024-04-17")] -alias nat_cast_lt_toNNReal := natCast_lt_toNNReal - -@[simp] -lemma toNNReal_le_ofNat {r : ℝ} {n : ℕ} [n.AtLeastTwo] : - r.toNNReal ≤ no_index (OfNat.ofNat n) ↔ r ≤ n := - toNNReal_le_natCast - -@[simp] -lemma ofNat_lt_toNNReal {r : ℝ} {n : ℕ} [n.AtLeastTwo] : - no_index (OfNat.ofNat n) < r.toNNReal ↔ n < r := - natCast_lt_toNNReal - -@[simp] -theorem toNNReal_eq_toNNReal_iff {r p : ℝ} (hr : 0 ≤ r) (hp : 0 ≤ p) : - toNNReal r = toNNReal p ↔ r = p := by simp [← coe_inj, coe_toNNReal, hr, hp] - -@[simp] -theorem toNNReal_lt_toNNReal_iff' {r p : ℝ} : Real.toNNReal r < Real.toNNReal p ↔ r < p ∧ 0 < p := - NNReal.coe_lt_coe.symm.trans max_lt_max_left_iff - -theorem toNNReal_lt_toNNReal_iff {r p : ℝ} (h : 0 < p) : - Real.toNNReal r < Real.toNNReal p ↔ r < p := - toNNReal_lt_toNNReal_iff'.trans (and_iff_left h) - -theorem lt_of_toNNReal_lt {r p : ℝ} (h : r.toNNReal < p.toNNReal) : r < p := - (Real.toNNReal_lt_toNNReal_iff <| Real.toNNReal_pos.1 (ne_bot_of_gt h).bot_lt).1 h - -theorem toNNReal_lt_toNNReal_iff_of_nonneg {r p : ℝ} (hr : 0 ≤ r) : - Real.toNNReal r < Real.toNNReal p ↔ r < p := - toNNReal_lt_toNNReal_iff'.trans ⟨And.left, fun h => ⟨h, lt_of_le_of_lt hr h⟩⟩ - -lemma toNNReal_le_toNNReal_iff' {r p : ℝ} : r.toNNReal ≤ p.toNNReal ↔ r ≤ p ∨ r ≤ 0 := by - simp_rw [← not_lt, toNNReal_lt_toNNReal_iff', not_and_or] - -lemma toNNReal_le_toNNReal_iff_of_pos {r p : ℝ} (hr : 0 < r) : r.toNNReal ≤ p.toNNReal ↔ r ≤ p := by - simp [toNNReal_le_toNNReal_iff', hr.not_le] - -@[simp] -lemma one_le_toNNReal {r : ℝ} : 1 ≤ r.toNNReal ↔ 1 ≤ r := by - simpa using toNNReal_le_toNNReal_iff_of_pos one_pos - -@[simp] -lemma toNNReal_lt_one {r : ℝ} : r.toNNReal < 1 ↔ r < 1 := by simp only [← not_le, one_le_toNNReal] - -@[simp] -lemma natCastle_toNNReal' {n : ℕ} {r : ℝ} : ↑n ≤ r.toNNReal ↔ n ≤ r ∨ n = 0 := by - simpa [n.cast_nonneg.le_iff_eq] using toNNReal_le_toNNReal_iff' (r := n) - -@[deprecated (since := "2024-04-17")] -alias nat_cast_le_toNNReal' := natCastle_toNNReal' - -@[simp] -lemma toNNReal_lt_natCast' {n : ℕ} {r : ℝ} : r.toNNReal < n ↔ r < n ∧ n ≠ 0 := by - simpa [pos_iff_ne_zero] using toNNReal_lt_toNNReal_iff' (r := r) (p := n) - -@[deprecated (since := "2024-04-17")] -alias toNNReal_lt_nat_cast' := toNNReal_lt_natCast' - -lemma natCast_le_toNNReal {n : ℕ} {r : ℝ} (hn : n ≠ 0) : ↑n ≤ r.toNNReal ↔ n ≤ r := by simp [hn] - -@[deprecated (since := "2024-04-17")] -alias nat_cast_le_toNNReal := natCast_le_toNNReal - -lemma toNNReal_lt_natCast {r : ℝ} {n : ℕ} (hn : n ≠ 0) : r.toNNReal < n ↔ r < n := by simp [hn] - -@[deprecated (since := "2024-04-17")] -alias toNNReal_lt_nat_cast := toNNReal_lt_natCast - -@[simp] -lemma toNNReal_lt_ofNat {r : ℝ} {n : ℕ} [n.AtLeastTwo] : - r.toNNReal < no_index (OfNat.ofNat n) ↔ r < OfNat.ofNat n := - toNNReal_lt_natCast (NeZero.ne n) - -@[simp] -lemma ofNat_le_toNNReal {n : ℕ} {r : ℝ} [n.AtLeastTwo] : - no_index (OfNat.ofNat n) ≤ r.toNNReal ↔ OfNat.ofNat n ≤ r := - natCast_le_toNNReal (NeZero.ne n) - -@[simp] -theorem toNNReal_add {r p : ℝ} (hr : 0 ≤ r) (hp : 0 ≤ p) : - Real.toNNReal (r + p) = Real.toNNReal r + Real.toNNReal p := - NNReal.eq <| by simp [hr, hp, add_nonneg] - -theorem toNNReal_add_toNNReal {r p : ℝ} (hr : 0 ≤ r) (hp : 0 ≤ p) : - Real.toNNReal r + Real.toNNReal p = Real.toNNReal (r + p) := - (Real.toNNReal_add hr hp).symm - -theorem toNNReal_le_toNNReal {r p : ℝ} (h : r ≤ p) : Real.toNNReal r ≤ Real.toNNReal p := - Real.toNNReal_mono h - -theorem toNNReal_add_le {r p : ℝ} : Real.toNNReal (r + p) ≤ Real.toNNReal r + Real.toNNReal p := - NNReal.coe_le_coe.1 <| max_le (add_le_add (le_max_left _ _) (le_max_left _ _)) NNReal.zero_le_coe - -theorem toNNReal_le_iff_le_coe {r : ℝ} {p : ℝ≥0} : toNNReal r ≤ p ↔ r ≤ ↑p := - NNReal.gi.gc r p - -theorem le_toNNReal_iff_coe_le {r : ℝ≥0} {p : ℝ} (hp : 0 ≤ p) : r ≤ Real.toNNReal p ↔ ↑r ≤ p := by - rw [← NNReal.coe_le_coe, Real.coe_toNNReal p hp] - -theorem le_toNNReal_iff_coe_le' {r : ℝ≥0} {p : ℝ} (hr : 0 < r) : r ≤ Real.toNNReal p ↔ ↑r ≤ p := - (le_or_lt 0 p).elim le_toNNReal_iff_coe_le fun hp => by - simp only [(hp.trans_le r.coe_nonneg).not_le, toNNReal_eq_zero.2 hp.le, hr.not_le] - -theorem toNNReal_lt_iff_lt_coe {r : ℝ} {p : ℝ≥0} (ha : 0 ≤ r) : Real.toNNReal r < p ↔ r < ↑p := by - rw [← NNReal.coe_lt_coe, Real.coe_toNNReal r ha] - -theorem lt_toNNReal_iff_coe_lt {r : ℝ≥0} {p : ℝ} : r < Real.toNNReal p ↔ ↑r < p := - lt_iff_lt_of_le_iff_le toNNReal_le_iff_le_coe - -theorem toNNReal_pow {x : ℝ} (hx : 0 ≤ x) (n : ℕ) : (x ^ n).toNNReal = x.toNNReal ^ n := by - rw [← coe_inj, NNReal.coe_pow, Real.coe_toNNReal _ (pow_nonneg hx _), - Real.coe_toNNReal x hx] - -theorem toNNReal_mul {p q : ℝ} (hp : 0 ≤ p) : - Real.toNNReal (p * q) = Real.toNNReal p * Real.toNNReal q := - NNReal.eq <| by simp [mul_max_of_nonneg, hp] - -end ToNNReal - -end Real - open Real -namespace NNReal - -section Mul - -theorem mul_eq_mul_left {a b c : ℝ≥0} (h : a ≠ 0) : a * b = a * c ↔ b = c := by - rw [mul_eq_mul_left_iff, or_iff_left h] - -end Mul - -section Pow - -theorem pow_antitone_exp {a : ℝ≥0} (m n : ℕ) (mn : m ≤ n) (a1 : a ≤ 1) : a ^ n ≤ a ^ m := - pow_le_pow_of_le_one (zero_le a) a1 mn - -nonrec theorem exists_pow_lt_of_lt_one {a b : ℝ≥0} (ha : 0 < a) (hb : b < 1) : - ∃ n : ℕ, b ^ n < a := by - simpa only [← coe_pow, NNReal.coe_lt_coe] using - exists_pow_lt_of_lt_one (NNReal.coe_pos.2 ha) (NNReal.coe_lt_coe.2 hb) - -nonrec theorem exists_mem_Ico_zpow {x : ℝ≥0} {y : ℝ≥0} (hx : x ≠ 0) (hy : 1 < y) : - ∃ n : ℤ, x ∈ Set.Ico (y ^ n) (y ^ (n + 1)) := - exists_mem_Ico_zpow (α := ℝ) hx.bot_lt hy - -nonrec theorem exists_mem_Ioc_zpow {x : ℝ≥0} {y : ℝ≥0} (hx : x ≠ 0) (hy : 1 < y) : - ∃ n : ℤ, x ∈ Set.Ioc (y ^ n) (y ^ (n + 1)) := - exists_mem_Ioc_zpow (α := ℝ) hx.bot_lt hy - -end Pow - section Sub /-! @@ -783,166 +104,17 @@ typeclass. For lemmas about subtraction and addition see lemmas about `OrderedSu `Mathlib.Algebra.Order.Sub.Basic`. See also `mul_tsub` and `tsub_mul`. -/ -theorem sub_def {r p : ℝ≥0} : r - p = Real.toNNReal (r - p) := - rfl - -theorem coe_sub_def {r p : ℝ≥0} : ↑(r - p) = max (r - p : ℝ) 0 := - rfl - -example : OrderedSub ℝ≥0 := by infer_instance - theorem sub_div (a b c : ℝ≥0) : (a - b) / c = a / c - b / c := tsub_div _ _ _ end Sub -section Inv - -@[simp] -theorem inv_le {r p : ℝ≥0} (h : r ≠ 0) : r⁻¹ ≤ p ↔ 1 ≤ r * p := by - rw [← mul_le_mul_left (pos_iff_ne_zero.2 h), mul_inv_cancel₀ h] - -theorem inv_le_of_le_mul {r p : ℝ≥0} (h : 1 ≤ r * p) : r⁻¹ ≤ p := by - by_cases r = 0 <;> simp [*, inv_le] - -@[simp] -theorem le_inv_iff_mul_le {r p : ℝ≥0} (h : p ≠ 0) : r ≤ p⁻¹ ↔ r * p ≤ 1 := by - rw [← mul_le_mul_left (pos_iff_ne_zero.2 h), mul_inv_cancel₀ h, mul_comm] - -@[simp] -theorem lt_inv_iff_mul_lt {r p : ℝ≥0} (h : p ≠ 0) : r < p⁻¹ ↔ r * p < 1 := by - rw [← mul_lt_mul_left (pos_iff_ne_zero.2 h), mul_inv_cancel₀ h, mul_comm] - -@[deprecated le_inv_mul_iff₀ (since := "2024-08-21")] -theorem mul_le_iff_le_inv {a b r : ℝ≥0} (hr : r ≠ 0) : r * a ≤ b ↔ a ≤ r⁻¹ * b := - (le_inv_mul_iff₀ (pos_iff_ne_zero.2 hr)).symm - -@[deprecated le_div_iff₀ (since := "2024-08-21")] -theorem le_div_iff_mul_le {a b r : ℝ≥0} (hr : r ≠ 0) : a ≤ b / r ↔ a * r ≤ b := - le_div_iff₀ (pos_iff_ne_zero.2 hr) - -@[deprecated div_le_iff₀ (since := "2024-08-21")] -protected lemma div_le_iff {a b r : ℝ≥0} (hr : r ≠ 0) : a / r ≤ b ↔ a ≤ b * r := - div_le_iff₀ (pos_iff_ne_zero.2 hr) - -@[deprecated div_le_iff₀' (since := "2024-08-21")] -protected lemma div_le_iff' {a b r : ℝ≥0} (hr : r ≠ 0) : a / r ≤ b ↔ a ≤ r * b := - div_le_iff₀' (pos_iff_ne_zero.2 hr) - -theorem div_le_of_le_mul {a b c : ℝ≥0} (h : a ≤ b * c) : a / c ≤ b := - if h0 : c = 0 then by simp [h0] else (div_le_iff₀ (pos_iff_ne_zero.2 h0)).2 h - -theorem div_le_of_le_mul' {a b c : ℝ≥0} (h : a ≤ b * c) : a / b ≤ c := - div_le_of_le_mul <| mul_comm b c ▸ h - -@[deprecated le_div_iff₀ (since := "2024-08-21")] -protected lemma le_div_iff {a b r : ℝ≥0} (hr : r ≠ 0) : a ≤ b / r ↔ a * r ≤ b := - le_div_iff₀ hr.bot_lt - -@[deprecated le_div_iff₀' (since := "2024-10-02")] -theorem le_div_iff' {a b r : ℝ≥0} (hr : r ≠ 0) : a ≤ b / r ↔ r * a ≤ b := le_div_iff₀' hr.bot_lt - -@[deprecated div_lt_iff₀ (since := "2024-10-02")] -theorem div_lt_iff {a b r : ℝ≥0} (hr : r ≠ 0) : a / r < b ↔ a < b * r := div_lt_iff₀ hr.bot_lt - -@[deprecated div_lt_iff₀' (since := "2024-10-02")] -theorem div_lt_iff' {a b r : ℝ≥0} (hr : r ≠ 0) : a / r < b ↔ a < r * b := div_lt_iff₀' hr.bot_lt - -@[deprecated lt_div_iff₀ (since := "2024-10-02")] -theorem lt_div_iff {a b r : ℝ≥0} (hr : r ≠ 0) : a < b / r ↔ a * r < b := lt_div_iff₀ hr.bot_lt - -@[deprecated lt_div_iff₀' (since := "2024-10-02")] -theorem lt_div_iff' {a b r : ℝ≥0} (hr : r ≠ 0) : a < b / r ↔ r * a < b := lt_div_iff₀' hr.bot_lt - -theorem mul_lt_of_lt_div {a b r : ℝ≥0} (h : a < b / r) : a * r < b := - (lt_div_iff₀ <| pos_iff_ne_zero.2 fun hr => False.elim <| by simp [hr] at h).1 h - -theorem div_le_div_left_of_le {a b c : ℝ≥0} (c0 : c ≠ 0) (cb : c ≤ b) : - a / b ≤ a / c := - div_le_div_of_nonneg_left (zero_le _) c0.bot_lt cb - -nonrec theorem div_le_div_left {a b c : ℝ≥0} (a0 : 0 < a) (b0 : 0 < b) (c0 : 0 < c) : - a / b ≤ a / c ↔ c ≤ b := - div_le_div_left a0 b0 c0 - -theorem le_of_forall_lt_one_mul_le {x y : ℝ≥0} (h : ∀ a < 1, a * x ≤ y) : x ≤ y := - le_of_forall_ge_of_dense fun a ha => by - have hx : x ≠ 0 := pos_iff_ne_zero.1 (lt_of_le_of_lt (zero_le _) ha) - have hx' : x⁻¹ ≠ 0 := by rwa [Ne, inv_eq_zero] - have : a * x⁻¹ < 1 := by rwa [← lt_inv_iff_mul_lt hx', inv_inv] - have : a * x⁻¹ * x ≤ y := h _ this - rwa [mul_assoc, inv_mul_cancel₀ hx, mul_one] at this - -nonrec theorem half_le_self (a : ℝ≥0) : a / 2 ≤ a := - half_le_self bot_le - -nonrec theorem half_lt_self {a : ℝ≥0} (h : a ≠ 0) : a / 2 < a := - half_lt_self h.bot_lt - -theorem div_lt_one_of_lt {a b : ℝ≥0} (h : a < b) : a / b < 1 := by - rwa [div_lt_iff₀ h.bot_lt, one_mul] - -theorem _root_.Real.toNNReal_inv {x : ℝ} : Real.toNNReal x⁻¹ = (Real.toNNReal x)⁻¹ := by - rcases le_total 0 x with hx | hx - · nth_rw 1 [← Real.coe_toNNReal x hx] - rw [← NNReal.coe_inv, Real.toNNReal_coe] - · rw [toNNReal_eq_zero.mpr hx, inv_zero, toNNReal_eq_zero.mpr (inv_nonpos.mpr hx)] - -theorem _root_.Real.toNNReal_div {x y : ℝ} (hx : 0 ≤ x) : - Real.toNNReal (x / y) = Real.toNNReal x / Real.toNNReal y := by - rw [div_eq_mul_inv, div_eq_mul_inv, ← Real.toNNReal_inv, ← Real.toNNReal_mul hx] - -theorem _root_.Real.toNNReal_div' {x y : ℝ} (hy : 0 ≤ y) : - Real.toNNReal (x / y) = Real.toNNReal x / Real.toNNReal y := by - rw [div_eq_inv_mul, div_eq_inv_mul, Real.toNNReal_mul (inv_nonneg.2 hy), Real.toNNReal_inv] - -theorem inv_lt_one_iff {x : ℝ≥0} (hx : x ≠ 0) : x⁻¹ < 1 ↔ 1 < x := by - rw [← one_div, div_lt_iff₀ hx.bot_lt, one_mul] - -@[deprecated zpow_pos (since := "2024-10-08")] -protected theorem zpow_pos {x : ℝ≥0} (hx : x ≠ 0) (n : ℤ) : 0 < x ^ n := zpow_pos hx.bot_lt _ - -theorem inv_lt_inv {x y : ℝ≥0} (hx : x ≠ 0) (h : x < y) : y⁻¹ < x⁻¹ := - inv_strictAnti₀ hx.bot_lt h - -end Inv - -@[simp] -theorem abs_eq (x : ℝ≥0) : |(x : ℝ)| = x := - abs_of_nonneg x.property - section Csupr open Set variable {ι : Sort*} {f : ι → ℝ≥0} -theorem le_toNNReal_of_coe_le {x : ℝ≥0} {y : ℝ} (h : ↑x ≤ y) : x ≤ y.toNNReal := - (le_toNNReal_iff_coe_le <| x.2.trans h).2 h - -nonrec theorem sSup_of_not_bddAbove {s : Set ℝ≥0} (hs : ¬BddAbove s) : SupSet.sSup s = 0 := by - rw [← bddAbove_coe] at hs - rw [← coe_inj, coe_sSup, NNReal.coe_zero] - exact sSup_of_not_bddAbove hs - -theorem iSup_of_not_bddAbove (hf : ¬BddAbove (range f)) : ⨆ i, f i = 0 := - sSup_of_not_bddAbove hf - -theorem iSup_empty [IsEmpty ι] (f : ι → ℝ≥0) : ⨆ i, f i = 0 := ciSup_of_empty f - -theorem iInf_empty [IsEmpty ι] (f : ι → ℝ≥0) : ⨅ i, f i = 0 := by - rw [_root_.iInf_of_isEmpty, sInf_empty] - -@[simp] lemma iSup_eq_zero (hf : BddAbove (range f)) : ⨆ i, f i = 0 ↔ ∀ i, f i = 0 := by - cases isEmpty_or_nonempty ι - · simp - · simp [← bot_eq_zero', ← le_bot_iff, ciSup_le_iff hf] - -@[simp] -theorem iInf_const_zero {α : Sort*} : ⨅ _ : α, (0 : ℝ≥0) = 0 := by - rw [← coe_inj, coe_iInf] - exact Real.iInf_const_zero - theorem iInf_mul (f : ι → ℝ≥0) (a : ℝ≥0) : iInf f * a = ⨅ i, f i * a := by rw [← coe_inj, NNReal.coe_mul, coe_iInf, coe_iInf] exact Real.iInf_mul_of_nonneg (NNReal.coe_nonneg _) _ @@ -993,122 +165,3 @@ theorem le_iInf_mul_iInf {a : ℝ≥0} {g h : ι → ℝ≥0} (H : ∀ i j, a end Csupr end NNReal - -namespace Set - -namespace OrdConnected - -variable {s : Set ℝ} {t : Set ℝ≥0} - -theorem preimage_coe_nnreal_real (h : s.OrdConnected) : ((↑) ⁻¹' s : Set ℝ≥0).OrdConnected := - h.preimage_mono NNReal.coe_mono - -theorem image_coe_nnreal_real (h : t.OrdConnected) : ((↑) '' t : Set ℝ).OrdConnected := - ⟨forall_mem_image.2 fun x hx => - forall_mem_image.2 fun _y hy z hz => ⟨⟨z, x.2.trans hz.1⟩, h.out hx hy hz, rfl⟩⟩ - --- Porting note (#11215): TODO: does it generalize to a `GaloisInsertion`? -theorem image_real_toNNReal (h : s.OrdConnected) : (Real.toNNReal '' s).OrdConnected := by - refine ⟨forall_mem_image.2 fun x hx => forall_mem_image.2 fun y hy z hz => ?_⟩ - rcases le_total y 0 with hy₀ | hy₀ - · rw [mem_Icc, Real.toNNReal_of_nonpos hy₀, nonpos_iff_eq_zero] at hz - exact ⟨y, hy, (toNNReal_of_nonpos hy₀).trans hz.2.symm⟩ - · lift y to ℝ≥0 using hy₀ - rw [toNNReal_coe] at hz - exact ⟨z, h.out hx hy ⟨toNNReal_le_iff_le_coe.1 hz.1, hz.2⟩, toNNReal_coe⟩ - -theorem preimage_real_toNNReal (h : t.OrdConnected) : (Real.toNNReal ⁻¹' t).OrdConnected := - h.preimage_mono Real.toNNReal_mono - -end OrdConnected - -end Set - -namespace Real - -/-- The absolute value on `ℝ` as a map to `ℝ≥0`. -/ --- Porting note (kmill): `pp_nodot` has no affect here --- unless RFC lean4#1910 leads to dot notation for CoeFun -@[pp_nodot] -def nnabs : ℝ →*₀ ℝ≥0 where - toFun x := ⟨|x|, abs_nonneg x⟩ - map_zero' := by ext; simp - map_one' := by ext; simp - map_mul' x y := by ext; simp [abs_mul] - -@[norm_cast, simp] -theorem coe_nnabs (x : ℝ) : (nnabs x : ℝ) = |x| := - rfl - -@[simp] -theorem nnabs_of_nonneg {x : ℝ} (h : 0 ≤ x) : nnabs x = toNNReal x := by - ext - rw [coe_toNNReal x h, coe_nnabs, abs_of_nonneg h] - -theorem nnabs_coe (x : ℝ≥0) : nnabs x = x := by simp - -theorem coe_toNNReal_le (x : ℝ) : (toNNReal x : ℝ) ≤ |x| := - max_le (le_abs_self _) (abs_nonneg _) - -@[simp] lemma toNNReal_abs (x : ℝ) : |x|.toNNReal = nnabs x := NNReal.coe_injective <| by simp - -theorem cast_natAbs_eq_nnabs_cast (n : ℤ) : (n.natAbs : ℝ≥0) = nnabs n := by - ext - rw [NNReal.coe_natCast, Int.cast_natAbs, Real.coe_nnabs, Int.cast_abs] - -end Real - -section StrictMono - -open NNReal - -variable {Γ₀ : Type*} [LinearOrderedCommGroupWithZero Γ₀] - -/-- If `Γ₀ˣ` is nontrivial and `f : Γ₀ →*₀ ℝ≥0` is strictly monotone, then for any positive - `r : ℝ≥0`, there exists `d : Γ₀ˣ` with `f d < r`. -/ -theorem NNReal.exists_lt_of_strictMono [h : Nontrivial Γ₀ˣ] {f : Γ₀ →*₀ ℝ≥0} (hf : StrictMono f) - {r : ℝ≥0} (hr : 0 < r) : ∃ d : Γ₀ˣ, f d < r := by - obtain ⟨g, hg1⟩ := (nontrivial_iff_exists_ne (1 : Γ₀ˣ)).mp h - set u : Γ₀ˣ := if g < 1 then g else g⁻¹ with hu - have hfu : f u < 1 := by - rw [hu] - split_ifs with hu1 - · rw [← map_one f]; exact hf hu1 - · have hfg0 : f g ≠ 0 := - fun h0 ↦ (Units.ne_zero g) ((map_eq_zero f).mp h0) - have hg1' : 1 < g := lt_of_le_of_ne (not_lt.mp hu1) hg1.symm - rw [Units.val_inv_eq_inv_val, map_inv₀, inv_lt_one_iff hfg0, ← map_one f] - exact hf hg1' - obtain ⟨n, hn⟩ := exists_pow_lt_of_lt_one hr hfu - use u ^ n - rwa [Units.val_pow_eq_pow_val, map_pow] - -/-- If `Γ₀ˣ` is nontrivial and `f : Γ₀ →*₀ ℝ≥0` is strictly monotone, then for any positive - real `r`, there exists `d : Γ₀ˣ` with `f d < r`. -/ -theorem Real.exists_lt_of_strictMono [h : Nontrivial Γ₀ˣ] {f : Γ₀ →*₀ ℝ≥0} (hf : StrictMono f) - {r : ℝ} (hr : 0 < r) : ∃ d : Γ₀ˣ, (f d : ℝ) < r := by - set s : NNReal := ⟨r, le_of_lt hr⟩ - have hs : 0 < s := hr - exact NNReal.exists_lt_of_strictMono hf hs - -end StrictMono - -namespace Mathlib.Meta.Positivity - -open Lean Meta Qq Function - -private alias ⟨_, nnreal_coe_pos⟩ := coe_pos - -/-- Extension for the `positivity` tactic: cast from `ℝ≥0` to `ℝ`. -/ -@[positivity NNReal.toReal _] -def evalNNRealtoReal : PositivityExt where eval {u α} _zα _pα e := do - match u, α, e with - | 0, ~q(ℝ), ~q(NNReal.toReal $a) => - let ra ← core q(inferInstance) q(inferInstance) a - assertInstancesCommute - match ra with - | .positive pa => pure (.positive q(nnreal_coe_pos $pa)) - | _ => pure (.nonnegative q(NNReal.coe_nonneg $a)) - | _, _, _ => throwError "not NNReal.toReal" - -end Mathlib.Meta.Positivity diff --git a/Mathlib/Data/NNReal/Defs.lean b/Mathlib/Data/NNReal/Defs.lean new file mode 100644 index 0000000000000..b220a37efdd08 --- /dev/null +++ b/Mathlib/Data/NNReal/Defs.lean @@ -0,0 +1,999 @@ +/- +Copyright (c) 2018 Johan Commelin. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johan Commelin +-/ +import Mathlib.Algebra.Algebra.Defs +import Mathlib.Algebra.Order.Module.OrderedSMul +import Mathlib.Data.Real.Archimedean + +/-! +# Nonnegative real numbers + +In this file we define `NNReal` (notation: `ℝ≥0`) to be the type of non-negative real numbers, +a.k.a. the interval `[0, ∞)`. We also define the following operations and structures on `ℝ≥0`: + +* the order on `ℝ≥0` is the restriction of the order on `ℝ`; these relations define a conditionally + complete linear order with a bottom element, `ConditionallyCompleteLinearOrderBot`; + +* `a + b` and `a * b` are the restrictions of addition and multiplication of real numbers to `ℝ≥0`; + these operations together with `0 = ⟨0, _⟩` and `1 = ⟨1, _⟩` turn `ℝ≥0` into a conditionally + complete linear ordered archimedean commutative semifield; we have no typeclass for this in + `mathlib` yet, so we define the following instances instead: + + - `LinearOrderedSemiring ℝ≥0`; + - `OrderedCommSemiring ℝ≥0`; + - `CanonicallyOrderedCommSemiring ℝ≥0`; + - `LinearOrderedCommGroupWithZero ℝ≥0`; + - `CanonicallyLinearOrderedAddCommMonoid ℝ≥0`; + - `Archimedean ℝ≥0`; + - `ConditionallyCompleteLinearOrderBot ℝ≥0`. + + These instances are derived from corresponding instances about the type `{x : α // 0 ≤ x}` in an + appropriate ordered field/ring/group/monoid `α`, see `Mathlib.Algebra.Order.Nonneg.OrderedRing`. + +* `Real.toNNReal x` is defined as `⟨max x 0, _⟩`, i.e. `↑(Real.toNNReal x) = x` when `0 ≤ x` and + `↑(Real.toNNReal x) = 0` otherwise. + +We also define an instance `CanLift ℝ ℝ≥0`. This instance can be used by the `lift` tactic to +replace `x : ℝ` and `hx : 0 ≤ x` in the proof context with `x : ℝ≥0` while replacing all occurrences +of `x` with `↑x`. This tactic also works for a function `f : α → ℝ` with a hypothesis +`hf : ∀ x, 0 ≤ f x`. + +## Notations + +This file defines `ℝ≥0` as a localized notation for `NNReal`. +-/ + +assert_not_exists Star + +open Function + +-- to ensure these instances are computable +/-- Nonnegative real numbers. -/ +def NNReal := { r : ℝ // 0 ≤ r } deriving + Zero, One, Semiring, StrictOrderedSemiring, CommMonoidWithZero, CommSemiring, + SemilatticeInf, SemilatticeSup, DistribLattice, OrderedCommSemiring, + CanonicallyOrderedCommSemiring, Inhabited + +namespace NNReal + +scoped notation "ℝ≥0" => NNReal + +instance instDenselyOrdered : DenselyOrdered ℝ≥0 := Nonneg.instDenselyOrdered +instance : OrderBot ℝ≥0 := inferInstance +instance instArchimedean : Archimedean ℝ≥0 := Nonneg.instArchimedean +instance instMulArchimedean : MulArchimedean ℝ≥0 := Nonneg.instMulArchimedean +noncomputable instance : Sub ℝ≥0 := Nonneg.sub +noncomputable instance : OrderedSub ℝ≥0 := Nonneg.orderedSub + +noncomputable instance : CanonicallyLinearOrderedSemifield ℝ≥0 := + Nonneg.canonicallyLinearOrderedSemifield + +/-- Coercion `ℝ≥0 → ℝ`. -/ +@[coe] def toReal : ℝ≥0 → ℝ := Subtype.val + +instance : Coe ℝ≥0 ℝ := ⟨toReal⟩ + +-- Simp lemma to put back `n.val` into the normal form given by the coercion. +@[simp] +theorem val_eq_coe (n : ℝ≥0) : n.val = n := + rfl + +instance canLift : CanLift ℝ ℝ≥0 toReal fun r => 0 ≤ r := + Subtype.canLift _ + +@[ext] protected theorem eq {n m : ℝ≥0} : (n : ℝ) = (m : ℝ) → n = m := + Subtype.eq + +theorem ne_iff {x y : ℝ≥0} : (x : ℝ) ≠ (y : ℝ) ↔ x ≠ y := + not_congr <| NNReal.eq_iff.symm + +protected theorem «forall» {p : ℝ≥0 → Prop} : + (∀ x : ℝ≥0, p x) ↔ ∀ (x : ℝ) (hx : 0 ≤ x), p ⟨x, hx⟩ := + Subtype.forall + +protected theorem «exists» {p : ℝ≥0 → Prop} : + (∃ x : ℝ≥0, p x) ↔ ∃ (x : ℝ) (hx : 0 ≤ x), p ⟨x, hx⟩ := + Subtype.exists + +/-- Reinterpret a real number `r` as a non-negative real number. Returns `0` if `r < 0`. -/ +noncomputable def _root_.Real.toNNReal (r : ℝ) : ℝ≥0 := + ⟨max r 0, le_max_right _ _⟩ + +theorem _root_.Real.coe_toNNReal (r : ℝ) (hr : 0 ≤ r) : (Real.toNNReal r : ℝ) = r := + max_eq_left hr + +theorem _root_.Real.toNNReal_of_nonneg {r : ℝ} (hr : 0 ≤ r) : r.toNNReal = ⟨r, hr⟩ := by + simp_rw [Real.toNNReal, max_eq_left hr] + +theorem _root_.Real.le_coe_toNNReal (r : ℝ) : r ≤ Real.toNNReal r := + le_max_left r 0 + +@[bound] theorem coe_nonneg (r : ℝ≥0) : (0 : ℝ) ≤ r := r.2 + +@[simp, norm_cast] theorem coe_mk (a : ℝ) (ha) : toReal ⟨a, ha⟩ = a := rfl + +example : Zero ℝ≥0 := by infer_instance + +example : One ℝ≥0 := by infer_instance + +example : Add ℝ≥0 := by infer_instance + +noncomputable example : Sub ℝ≥0 := by infer_instance + +example : Mul ℝ≥0 := by infer_instance + +noncomputable example : Inv ℝ≥0 := by infer_instance + +noncomputable example : Div ℝ≥0 := by infer_instance + +example : LE ℝ≥0 := by infer_instance + +example : Bot ℝ≥0 := by infer_instance + +example : Inhabited ℝ≥0 := by infer_instance + +example : Nontrivial ℝ≥0 := by infer_instance + +protected theorem coe_injective : Injective ((↑) : ℝ≥0 → ℝ) := Subtype.coe_injective + +@[simp, norm_cast] lemma coe_inj {r₁ r₂ : ℝ≥0} : (r₁ : ℝ) = r₂ ↔ r₁ = r₂ := + NNReal.coe_injective.eq_iff + +@[deprecated (since := "2024-02-03")] protected alias coe_eq := coe_inj + +@[simp, norm_cast] lemma coe_zero : ((0 : ℝ≥0) : ℝ) = 0 := rfl + +@[simp, norm_cast] lemma coe_one : ((1 : ℝ≥0) : ℝ) = 1 := rfl + +@[simp] lemma mk_zero : (⟨0, le_rfl⟩ : ℝ≥0) = 0 := rfl +@[simp] lemma mk_one : (⟨1, zero_le_one⟩ : ℝ≥0) = 1 := rfl + +@[simp, norm_cast] +protected theorem coe_add (r₁ r₂ : ℝ≥0) : ((r₁ + r₂ : ℝ≥0) : ℝ) = r₁ + r₂ := + rfl + +@[simp, norm_cast] +protected theorem coe_mul (r₁ r₂ : ℝ≥0) : ((r₁ * r₂ : ℝ≥0) : ℝ) = r₁ * r₂ := + rfl + +@[simp, norm_cast] +protected theorem coe_inv (r : ℝ≥0) : ((r⁻¹ : ℝ≥0) : ℝ) = (r : ℝ)⁻¹ := + rfl + +@[simp, norm_cast] +protected theorem coe_div (r₁ r₂ : ℝ≥0) : ((r₁ / r₂ : ℝ≥0) : ℝ) = (r₁ : ℝ) / r₂ := + rfl + +protected theorem coe_two : ((2 : ℝ≥0) : ℝ) = 2 := rfl + +@[simp, norm_cast] +protected theorem coe_sub {r₁ r₂ : ℝ≥0} (h : r₂ ≤ r₁) : ((r₁ - r₂ : ℝ≥0) : ℝ) = ↑r₁ - ↑r₂ := + max_eq_left <| le_sub_comm.2 <| by simp [show (r₂ : ℝ) ≤ r₁ from h] + +variable {r r₁ r₂ : ℝ≥0} {x y : ℝ} + +@[simp, norm_cast] lemma coe_eq_zero : (r : ℝ) = 0 ↔ r = 0 := by rw [← coe_zero, coe_inj] + +@[simp, norm_cast] lemma coe_eq_one : (r : ℝ) = 1 ↔ r = 1 := by rw [← coe_one, coe_inj] + +@[norm_cast] lemma coe_ne_zero : (r : ℝ) ≠ 0 ↔ r ≠ 0 := coe_eq_zero.not + +@[norm_cast] lemma coe_ne_one : (r : ℝ) ≠ 1 ↔ r ≠ 1 := coe_eq_one.not + +example : CommSemiring ℝ≥0 := by infer_instance + +/-- Coercion `ℝ≥0 → ℝ` as a `RingHom`. + +Porting note (#11215): TODO: what if we define `Coe ℝ≥0 ℝ` using this function? -/ +def toRealHom : ℝ≥0 →+* ℝ where + toFun := (↑) + map_one' := NNReal.coe_one + map_mul' := NNReal.coe_mul + map_zero' := NNReal.coe_zero + map_add' := NNReal.coe_add + +@[simp] theorem coe_toRealHom : ⇑toRealHom = toReal := rfl + +section Actions + +/-- A `MulAction` over `ℝ` restricts to a `MulAction` over `ℝ≥0`. -/ +instance {M : Type*} [MulAction ℝ M] : MulAction ℝ≥0 M := + MulAction.compHom M toRealHom.toMonoidHom + +theorem smul_def {M : Type*} [MulAction ℝ M] (c : ℝ≥0) (x : M) : c • x = (c : ℝ) • x := + rfl + +instance {M N : Type*} [MulAction ℝ M] [MulAction ℝ N] [SMul M N] [IsScalarTower ℝ M N] : + IsScalarTower ℝ≥0 M N where smul_assoc r := (smul_assoc (r : ℝ) : _) + +instance smulCommClass_left {M N : Type*} [MulAction ℝ N] [SMul M N] [SMulCommClass ℝ M N] : + SMulCommClass ℝ≥0 M N where smul_comm r := (smul_comm (r : ℝ) : _) + +instance smulCommClass_right {M N : Type*} [MulAction ℝ N] [SMul M N] [SMulCommClass M ℝ N] : + SMulCommClass M ℝ≥0 N where smul_comm m r := (smul_comm m (r : ℝ) : _) + +/-- A `DistribMulAction` over `ℝ` restricts to a `DistribMulAction` over `ℝ≥0`. -/ +instance {M : Type*} [AddMonoid M] [DistribMulAction ℝ M] : DistribMulAction ℝ≥0 M := + DistribMulAction.compHom M toRealHom.toMonoidHom + +/-- A `Module` over `ℝ` restricts to a `Module` over `ℝ≥0`. -/ +instance {M : Type*} [AddCommMonoid M] [Module ℝ M] : Module ℝ≥0 M := + Module.compHom M toRealHom + +-- Porting note (#11215): TODO: after this line, `↑` uses `Algebra.cast` instead of `toReal` +/-- An `Algebra` over `ℝ` restricts to an `Algebra` over `ℝ≥0`. -/ +instance {A : Type*} [Semiring A] [Algebra ℝ A] : Algebra ℝ≥0 A where + smul := (· • ·) + commutes' r x := by simp [Algebra.commutes] + smul_def' r x := by simp [← Algebra.smul_def (r : ℝ) x, smul_def] + toRingHom := (algebraMap ℝ A).comp (toRealHom : ℝ≥0 →+* ℝ) + +-- verify that the above produces instances we might care about +example : Algebra ℝ≥0 ℝ := by infer_instance + +example : DistribMulAction ℝ≥0ˣ ℝ := by infer_instance + +end Actions + +example : MonoidWithZero ℝ≥0 := by infer_instance + +example : CommMonoidWithZero ℝ≥0 := by infer_instance + +noncomputable example : CommGroupWithZero ℝ≥0 := by infer_instance + +@[simp, norm_cast] +theorem coe_pow (r : ℝ≥0) (n : ℕ) : ((r ^ n : ℝ≥0) : ℝ) = (r : ℝ) ^ n := rfl + +@[simp, norm_cast] +theorem coe_zpow (r : ℝ≥0) (n : ℤ) : ((r ^ n : ℝ≥0) : ℝ) = (r : ℝ) ^ n := rfl + +variable {ι : Type*} {f : ι → ℝ} + +@[simp, norm_cast] lemma coe_nsmul (r : ℝ≥0) (n : ℕ) : ↑(n • r) = n • (r : ℝ) := rfl +@[simp, norm_cast] lemma coe_nnqsmul (q : ℚ≥0) (x : ℝ≥0) : ↑(q • x) = (q • x : ℝ) := rfl + +@[simp, norm_cast] +protected theorem coe_natCast (n : ℕ) : (↑(↑n : ℝ≥0) : ℝ) = n := + map_natCast toRealHom n + +@[deprecated (since := "2024-04-17")] +alias coe_nat_cast := NNReal.coe_natCast + +-- See note [no_index around OfNat.ofNat] +@[simp, norm_cast] +protected theorem coe_ofNat (n : ℕ) [n.AtLeastTwo] : + (no_index (OfNat.ofNat n : ℝ≥0) : ℝ) = OfNat.ofNat n := + rfl + +@[simp, norm_cast] +protected theorem coe_ofScientific (m : ℕ) (s : Bool) (e : ℕ) : + ↑(OfScientific.ofScientific m s e : ℝ≥0) = (OfScientific.ofScientific m s e : ℝ) := + rfl + +@[simp, norm_cast] +lemma algebraMap_eq_coe : (algebraMap ℝ≥0 ℝ : ℝ≥0 → ℝ) = (↑) := rfl + +noncomputable example : LinearOrder ℝ≥0 := by infer_instance + +@[simp, norm_cast] lemma coe_le_coe : (r₁ : ℝ) ≤ r₂ ↔ r₁ ≤ r₂ := Iff.rfl + +@[simp, norm_cast] lemma coe_lt_coe : (r₁ : ℝ) < r₂ ↔ r₁ < r₂ := Iff.rfl + +@[bound] private alias ⟨_, Bound.coe_lt_coe_of_lt⟩ := coe_lt_coe + +@[simp, norm_cast] lemma coe_pos : (0 : ℝ) < r ↔ 0 < r := Iff.rfl + +@[bound] private alias ⟨_, Bound.coe_pos_of_pos⟩ := coe_pos + +@[simp, norm_cast] lemma one_le_coe : 1 ≤ (r : ℝ) ↔ 1 ≤ r := by rw [← coe_le_coe, coe_one] +@[simp, norm_cast] lemma one_lt_coe : 1 < (r : ℝ) ↔ 1 < r := by rw [← coe_lt_coe, coe_one] +@[simp, norm_cast] lemma coe_le_one : (r : ℝ) ≤ 1 ↔ r ≤ 1 := by rw [← coe_le_coe, coe_one] +@[simp, norm_cast] lemma coe_lt_one : (r : ℝ) < 1 ↔ r < 1 := by rw [← coe_lt_coe, coe_one] + +@[mono] lemma coe_mono : Monotone ((↑) : ℝ≥0 → ℝ) := fun _ _ => NNReal.coe_le_coe.2 + +/-- Alias for the use of `gcongr` -/ +@[gcongr] alias ⟨_, GCongr.toReal_le_toReal⟩ := coe_le_coe + +protected theorem _root_.Real.toNNReal_mono : Monotone Real.toNNReal := fun _ _ h => + max_le_max h (le_refl 0) + +@[simp] +theorem _root_.Real.toNNReal_coe {r : ℝ≥0} : Real.toNNReal r = r := + NNReal.eq <| max_eq_left r.2 + +@[simp] +theorem mk_natCast (n : ℕ) : @Eq ℝ≥0 (⟨(n : ℝ), n.cast_nonneg⟩ : ℝ≥0) n := + NNReal.eq (NNReal.coe_natCast n).symm + +@[deprecated (since := "2024-04-05")] alias mk_coe_nat := mk_natCast + +-- Porting note: place this in the `Real` namespace +@[simp] +theorem toNNReal_coe_nat (n : ℕ) : Real.toNNReal n = n := + NNReal.eq <| by simp [Real.coe_toNNReal] + +-- See note [no_index around OfNat.ofNat] +@[simp] +theorem _root_.Real.toNNReal_ofNat (n : ℕ) [n.AtLeastTwo] : + Real.toNNReal (no_index (OfNat.ofNat n)) = OfNat.ofNat n := + toNNReal_coe_nat n + +/-- `Real.toNNReal` and `NNReal.toReal : ℝ≥0 → ℝ` form a Galois insertion. -/ +noncomputable def gi : GaloisInsertion Real.toNNReal (↑) := + GaloisInsertion.monotoneIntro NNReal.coe_mono Real.toNNReal_mono Real.le_coe_toNNReal fun _ => + Real.toNNReal_coe + +-- note that anything involving the (decidability of the) linear order, +-- will be noncomputable, everything else should not be. +example : OrderBot ℝ≥0 := by infer_instance + +example : PartialOrder ℝ≥0 := by infer_instance + +noncomputable example : CanonicallyLinearOrderedAddCommMonoid ℝ≥0 := by infer_instance + +noncomputable example : LinearOrderedAddCommMonoid ℝ≥0 := by infer_instance + +example : DistribLattice ℝ≥0 := by infer_instance + +example : SemilatticeInf ℝ≥0 := by infer_instance + +example : SemilatticeSup ℝ≥0 := by infer_instance + +noncomputable example : LinearOrderedSemiring ℝ≥0 := by infer_instance + +example : OrderedCommSemiring ℝ≥0 := by infer_instance + +noncomputable example : LinearOrderedCommMonoid ℝ≥0 := by infer_instance + +noncomputable example : LinearOrderedCommMonoidWithZero ℝ≥0 := by infer_instance + +noncomputable example : LinearOrderedCommGroupWithZero ℝ≥0 := by infer_instance + +example : CanonicallyOrderedCommSemiring ℝ≥0 := by infer_instance + +example : DenselyOrdered ℝ≥0 := by infer_instance + +example : NoMaxOrder ℝ≥0 := by infer_instance + +instance instPosSMulStrictMono {α} [Preorder α] [MulAction ℝ α] [PosSMulStrictMono ℝ α] : + PosSMulStrictMono ℝ≥0 α where + elim _r hr _a₁ _a₂ ha := (smul_lt_smul_of_pos_left ha (coe_pos.2 hr):) + +instance instSMulPosStrictMono {α} [Zero α] [Preorder α] [MulAction ℝ α] [SMulPosStrictMono ℝ α] : + SMulPosStrictMono ℝ≥0 α where + elim _a ha _r₁ _r₂ hr := (smul_lt_smul_of_pos_right (coe_lt_coe.2 hr) ha :) + +/-- If `a` is a nonnegative real number, then the closed interval `[0, a]` in `ℝ` is order +isomorphic to the interval `Set.Iic a`. -/ +-- Porting note (#11215): TODO: restore once `simps` supports `ℝ≥0` @[simps!? apply_coe_coe] +def orderIsoIccZeroCoe (a : ℝ≥0) : Set.Icc (0 : ℝ) a ≃o Set.Iic a where + toEquiv := Equiv.Set.sep (Set.Ici 0) fun x : ℝ => x ≤ a + map_rel_iff' := Iff.rfl + +@[simp] +theorem orderIsoIccZeroCoe_apply_coe_coe (a : ℝ≥0) (b : Set.Icc (0 : ℝ) a) : + (orderIsoIccZeroCoe a b : ℝ) = b := + rfl + +@[simp] +theorem orderIsoIccZeroCoe_symm_apply_coe (a : ℝ≥0) (b : Set.Iic a) : + ((orderIsoIccZeroCoe a).symm b : ℝ) = b := + rfl + +-- note we need the `@` to make the `Membership.mem` have a sensible type +theorem coe_image {s : Set ℝ≥0} : + (↑) '' s = { x : ℝ | ∃ h : 0 ≤ x, @Membership.mem ℝ≥0 _ _ s ⟨x, h⟩ } := + Subtype.coe_image + +theorem bddAbove_coe {s : Set ℝ≥0} : BddAbove (((↑) : ℝ≥0 → ℝ) '' s) ↔ BddAbove s := + Iff.intro + (fun ⟨b, hb⟩ => + ⟨Real.toNNReal b, fun ⟨y, _⟩ hys => + show y ≤ max b 0 from le_max_of_le_left <| hb <| Set.mem_image_of_mem _ hys⟩) + fun ⟨b, hb⟩ => ⟨b, fun _ ⟨_, hx, eq⟩ => eq ▸ hb hx⟩ + +theorem bddBelow_coe (s : Set ℝ≥0) : BddBelow (((↑) : ℝ≥0 → ℝ) '' s) := + ⟨0, fun _ ⟨q, _, eq⟩ => eq ▸ q.2⟩ + +noncomputable instance : ConditionallyCompleteLinearOrderBot ℝ≥0 := + Nonneg.conditionallyCompleteLinearOrderBot 0 + +@[norm_cast] +theorem coe_sSup (s : Set ℝ≥0) : (↑(sSup s) : ℝ) = sSup (((↑) : ℝ≥0 → ℝ) '' s) := by + rcases Set.eq_empty_or_nonempty s with rfl|hs + · simp + by_cases H : BddAbove s + · have A : sSup (Subtype.val '' s) ∈ Set.Ici 0 := by + apply Real.sSup_nonneg + rintro - ⟨y, -, rfl⟩ + exact y.2 + exact (@subset_sSup_of_within ℝ (Set.Ici (0 : ℝ)) _ _ (_) s hs H A).symm + · simp only [csSup_of_not_bddAbove H, csSup_empty, bot_eq_zero', NNReal.coe_zero] + apply (Real.sSup_of_not_bddAbove ?_).symm + contrapose! H + exact bddAbove_coe.1 H + +@[simp, norm_cast] -- Porting note: add `simp` +theorem coe_iSup {ι : Sort*} (s : ι → ℝ≥0) : (↑(⨆ i, s i) : ℝ) = ⨆ i, ↑(s i) := by + rw [iSup, iSup, coe_sSup, ← Set.range_comp]; rfl + +@[norm_cast] +theorem coe_sInf (s : Set ℝ≥0) : (↑(sInf s) : ℝ) = sInf (((↑) : ℝ≥0 → ℝ) '' s) := by + rcases Set.eq_empty_or_nonempty s with rfl|hs + · simp only [Set.image_empty, Real.sInf_empty, coe_eq_zero] + exact @subset_sInf_emptyset ℝ (Set.Ici (0 : ℝ)) _ _ (_) + have A : sInf (Subtype.val '' s) ∈ Set.Ici 0 := by + apply Real.sInf_nonneg + rintro - ⟨y, -, rfl⟩ + exact y.2 + exact (@subset_sInf_of_within ℝ (Set.Ici (0 : ℝ)) _ _ (_) s hs (OrderBot.bddBelow s) A).symm + +@[simp] +theorem sInf_empty : sInf (∅ : Set ℝ≥0) = 0 := by + rw [← coe_eq_zero, coe_sInf, Set.image_empty, Real.sInf_empty] + +@[norm_cast] +theorem coe_iInf {ι : Sort*} (s : ι → ℝ≥0) : (↑(⨅ i, s i) : ℝ) = ⨅ i, ↑(s i) := by + rw [iInf, iInf, coe_sInf, ← Set.range_comp]; rfl + +-- Short-circuit instance search +instance instCovariantClassAddLE : CovariantClass ℝ≥0 ℝ≥0 (· + ·) (· ≤ ·) := inferInstance +instance instContravariantClassAddLT : ContravariantClass ℝ≥0 ℝ≥0 (· + ·) (· < ·) := inferInstance +instance instCovariantClassMulLE : CovariantClass ℝ≥0 ℝ≥0 (· * ·) (· ≤ ·) := inferInstance + +@[deprecated le_of_forall_pos_le_add (since := "2024-10-17")] +protected theorem le_of_forall_pos_le_add {a b : ℝ≥0} (h : ∀ ε, 0 < ε → a ≤ b + ε) : a ≤ b := + le_of_forall_pos_le_add h + +theorem lt_iff_exists_rat_btwn (a b : ℝ≥0) : + a < b ↔ ∃ q : ℚ, 0 ≤ q ∧ a < Real.toNNReal q ∧ Real.toNNReal q < b := + Iff.intro + (fun h : (↑a : ℝ) < (↑b : ℝ) => + let ⟨q, haq, hqb⟩ := exists_rat_btwn h + have : 0 ≤ (q : ℝ) := le_trans a.2 <| le_of_lt haq + ⟨q, Rat.cast_nonneg.1 this, by + simp [Real.coe_toNNReal _ this, NNReal.coe_lt_coe.symm, haq, hqb]⟩) + fun ⟨_, _, haq, hqb⟩ => lt_trans haq hqb + +theorem bot_eq_zero : (⊥ : ℝ≥0) = 0 := rfl + +theorem mul_sup (a b c : ℝ≥0) : a * (b ⊔ c) = a * b ⊔ a * c := + mul_max_of_nonneg _ _ <| zero_le a + +theorem sup_mul (a b c : ℝ≥0) : (a ⊔ b) * c = a * c ⊔ b * c := + max_mul_of_nonneg _ _ <| zero_le c + +@[simp, norm_cast] +theorem coe_max (x y : ℝ≥0) : ((max x y : ℝ≥0) : ℝ) = max (x : ℝ) (y : ℝ) := + NNReal.coe_mono.map_max + +@[simp, norm_cast] +theorem coe_min (x y : ℝ≥0) : ((min x y : ℝ≥0) : ℝ) = min (x : ℝ) (y : ℝ) := + NNReal.coe_mono.map_min + +@[simp] +theorem zero_le_coe {q : ℝ≥0} : 0 ≤ (q : ℝ) := + q.2 + +instance instOrderedSMul {M : Type*} [OrderedAddCommMonoid M] [Module ℝ M] [OrderedSMul ℝ M] : + OrderedSMul ℝ≥0 M where + smul_lt_smul_of_pos hab hc := (smul_lt_smul_of_pos_left hab (NNReal.coe_pos.2 hc) : _) + lt_of_smul_lt_smul_of_pos {_ _ c} hab _ := + lt_of_smul_lt_smul_of_nonneg_left (by exact hab) (NNReal.coe_nonneg c) + +end NNReal + +open NNReal + +namespace Real + +section ToNNReal + +@[simp] +theorem coe_toNNReal' (r : ℝ) : (Real.toNNReal r : ℝ) = max r 0 := + rfl + +@[simp] +theorem toNNReal_zero : Real.toNNReal 0 = 0 := NNReal.eq <| coe_toNNReal _ le_rfl + +@[simp] +theorem toNNReal_one : Real.toNNReal 1 = 1 := NNReal.eq <| coe_toNNReal _ zero_le_one + +@[simp] +theorem toNNReal_pos {r : ℝ} : 0 < Real.toNNReal r ↔ 0 < r := by + simp [← NNReal.coe_lt_coe, lt_irrefl] + +@[simp] +theorem toNNReal_eq_zero {r : ℝ} : Real.toNNReal r = 0 ↔ r ≤ 0 := by + simpa [-toNNReal_pos] using not_iff_not.2 (@toNNReal_pos r) + +theorem toNNReal_of_nonpos {r : ℝ} : r ≤ 0 → Real.toNNReal r = 0 := + toNNReal_eq_zero.2 + +lemma toNNReal_eq_iff_eq_coe {r : ℝ} {p : ℝ≥0} (hp : p ≠ 0) : r.toNNReal = p ↔ r = p := + ⟨fun h ↦ h ▸ (coe_toNNReal _ <| not_lt.1 fun hlt ↦ hp <| h ▸ toNNReal_of_nonpos hlt.le).symm, + fun h ↦ h.symm ▸ toNNReal_coe⟩ + +@[simp] +lemma toNNReal_eq_one {r : ℝ} : r.toNNReal = 1 ↔ r = 1 := toNNReal_eq_iff_eq_coe one_ne_zero + +@[simp] +lemma toNNReal_eq_natCast {r : ℝ} {n : ℕ} (hn : n ≠ 0) : r.toNNReal = n ↔ r = n := + mod_cast toNNReal_eq_iff_eq_coe <| Nat.cast_ne_zero.2 hn + +@[deprecated (since := "2024-04-17")] +alias toNNReal_eq_nat_cast := toNNReal_eq_natCast + +@[simp] +lemma toNNReal_eq_ofNat {r : ℝ} {n : ℕ} [n.AtLeastTwo] : + r.toNNReal = no_index (OfNat.ofNat n) ↔ r = OfNat.ofNat n := + toNNReal_eq_natCast (NeZero.ne n) + +@[simp] +theorem toNNReal_le_toNNReal_iff {r p : ℝ} (hp : 0 ≤ p) : + toNNReal r ≤ toNNReal p ↔ r ≤ p := by simp [← NNReal.coe_le_coe, hp] + +@[simp] +lemma toNNReal_le_one {r : ℝ} : r.toNNReal ≤ 1 ↔ r ≤ 1 := by + simpa using toNNReal_le_toNNReal_iff zero_le_one + +@[simp] +lemma one_lt_toNNReal {r : ℝ} : 1 < r.toNNReal ↔ 1 < r := by + simpa only [not_le] using toNNReal_le_one.not + +@[simp] +lemma toNNReal_le_natCast {r : ℝ} {n : ℕ} : r.toNNReal ≤ n ↔ r ≤ n := by + simpa using toNNReal_le_toNNReal_iff n.cast_nonneg + +@[deprecated (since := "2024-04-17")] +alias toNNReal_le_nat_cast := toNNReal_le_natCast + +@[simp] +lemma natCast_lt_toNNReal {r : ℝ} {n : ℕ} : n < r.toNNReal ↔ n < r := by + simpa only [not_le] using toNNReal_le_natCast.not + +@[deprecated (since := "2024-04-17")] +alias nat_cast_lt_toNNReal := natCast_lt_toNNReal + +@[simp] +lemma toNNReal_le_ofNat {r : ℝ} {n : ℕ} [n.AtLeastTwo] : + r.toNNReal ≤ no_index (OfNat.ofNat n) ↔ r ≤ n := + toNNReal_le_natCast + +@[simp] +lemma ofNat_lt_toNNReal {r : ℝ} {n : ℕ} [n.AtLeastTwo] : + no_index (OfNat.ofNat n) < r.toNNReal ↔ n < r := + natCast_lt_toNNReal + +@[simp] +theorem toNNReal_eq_toNNReal_iff {r p : ℝ} (hr : 0 ≤ r) (hp : 0 ≤ p) : + toNNReal r = toNNReal p ↔ r = p := by simp [← coe_inj, coe_toNNReal, hr, hp] + +@[simp] +theorem toNNReal_lt_toNNReal_iff' {r p : ℝ} : Real.toNNReal r < Real.toNNReal p ↔ r < p ∧ 0 < p := + NNReal.coe_lt_coe.symm.trans max_lt_max_left_iff + +theorem toNNReal_lt_toNNReal_iff {r p : ℝ} (h : 0 < p) : + Real.toNNReal r < Real.toNNReal p ↔ r < p := + toNNReal_lt_toNNReal_iff'.trans (and_iff_left h) + +theorem lt_of_toNNReal_lt {r p : ℝ} (h : r.toNNReal < p.toNNReal) : r < p := + (Real.toNNReal_lt_toNNReal_iff <| Real.toNNReal_pos.1 (ne_bot_of_gt h).bot_lt).1 h + +theorem toNNReal_lt_toNNReal_iff_of_nonneg {r p : ℝ} (hr : 0 ≤ r) : + Real.toNNReal r < Real.toNNReal p ↔ r < p := + toNNReal_lt_toNNReal_iff'.trans ⟨And.left, fun h => ⟨h, lt_of_le_of_lt hr h⟩⟩ + +lemma toNNReal_le_toNNReal_iff' {r p : ℝ} : r.toNNReal ≤ p.toNNReal ↔ r ≤ p ∨ r ≤ 0 := by + simp_rw [← not_lt, toNNReal_lt_toNNReal_iff', not_and_or] + +lemma toNNReal_le_toNNReal_iff_of_pos {r p : ℝ} (hr : 0 < r) : r.toNNReal ≤ p.toNNReal ↔ r ≤ p := by + simp [toNNReal_le_toNNReal_iff', hr.not_le] + +@[simp] +lemma one_le_toNNReal {r : ℝ} : 1 ≤ r.toNNReal ↔ 1 ≤ r := by + simpa using toNNReal_le_toNNReal_iff_of_pos one_pos + +@[simp] +lemma toNNReal_lt_one {r : ℝ} : r.toNNReal < 1 ↔ r < 1 := by simp only [← not_le, one_le_toNNReal] + +@[simp] +lemma natCastle_toNNReal' {n : ℕ} {r : ℝ} : ↑n ≤ r.toNNReal ↔ n ≤ r ∨ n = 0 := by + simpa [n.cast_nonneg.le_iff_eq] using toNNReal_le_toNNReal_iff' (r := n) + +@[deprecated (since := "2024-04-17")] +alias nat_cast_le_toNNReal' := natCastle_toNNReal' + +@[simp] +lemma toNNReal_lt_natCast' {n : ℕ} {r : ℝ} : r.toNNReal < n ↔ r < n ∧ n ≠ 0 := by + simpa [pos_iff_ne_zero] using toNNReal_lt_toNNReal_iff' (r := r) (p := n) + +@[deprecated (since := "2024-04-17")] +alias toNNReal_lt_nat_cast' := toNNReal_lt_natCast' + +lemma natCast_le_toNNReal {n : ℕ} {r : ℝ} (hn : n ≠ 0) : ↑n ≤ r.toNNReal ↔ n ≤ r := by simp [hn] + +@[deprecated (since := "2024-04-17")] +alias nat_cast_le_toNNReal := natCast_le_toNNReal + +lemma toNNReal_lt_natCast {r : ℝ} {n : ℕ} (hn : n ≠ 0) : r.toNNReal < n ↔ r < n := by simp [hn] + +@[deprecated (since := "2024-04-17")] +alias toNNReal_lt_nat_cast := toNNReal_lt_natCast + +@[simp] +lemma toNNReal_lt_ofNat {r : ℝ} {n : ℕ} [n.AtLeastTwo] : + r.toNNReal < no_index (OfNat.ofNat n) ↔ r < OfNat.ofNat n := + toNNReal_lt_natCast (NeZero.ne n) + +@[simp] +lemma ofNat_le_toNNReal {n : ℕ} {r : ℝ} [n.AtLeastTwo] : + no_index (OfNat.ofNat n) ≤ r.toNNReal ↔ OfNat.ofNat n ≤ r := + natCast_le_toNNReal (NeZero.ne n) + +@[simp] +theorem toNNReal_add {r p : ℝ} (hr : 0 ≤ r) (hp : 0 ≤ p) : + Real.toNNReal (r + p) = Real.toNNReal r + Real.toNNReal p := + NNReal.eq <| by simp [hr, hp, add_nonneg] + +theorem toNNReal_add_toNNReal {r p : ℝ} (hr : 0 ≤ r) (hp : 0 ≤ p) : + Real.toNNReal r + Real.toNNReal p = Real.toNNReal (r + p) := + (Real.toNNReal_add hr hp).symm + +theorem toNNReal_le_toNNReal {r p : ℝ} (h : r ≤ p) : Real.toNNReal r ≤ Real.toNNReal p := + Real.toNNReal_mono h + +theorem toNNReal_add_le {r p : ℝ} : Real.toNNReal (r + p) ≤ Real.toNNReal r + Real.toNNReal p := + NNReal.coe_le_coe.1 <| max_le (add_le_add (le_max_left _ _) (le_max_left _ _)) NNReal.zero_le_coe + +theorem toNNReal_le_iff_le_coe {r : ℝ} {p : ℝ≥0} : toNNReal r ≤ p ↔ r ≤ ↑p := + NNReal.gi.gc r p + +theorem le_toNNReal_iff_coe_le {r : ℝ≥0} {p : ℝ} (hp : 0 ≤ p) : r ≤ Real.toNNReal p ↔ ↑r ≤ p := by + rw [← NNReal.coe_le_coe, Real.coe_toNNReal p hp] + +theorem le_toNNReal_iff_coe_le' {r : ℝ≥0} {p : ℝ} (hr : 0 < r) : r ≤ Real.toNNReal p ↔ ↑r ≤ p := + (le_or_lt 0 p).elim le_toNNReal_iff_coe_le fun hp => by + simp only [(hp.trans_le r.coe_nonneg).not_le, toNNReal_eq_zero.2 hp.le, hr.not_le] + +theorem toNNReal_lt_iff_lt_coe {r : ℝ} {p : ℝ≥0} (ha : 0 ≤ r) : Real.toNNReal r < p ↔ r < ↑p := by + rw [← NNReal.coe_lt_coe, Real.coe_toNNReal r ha] + +theorem lt_toNNReal_iff_coe_lt {r : ℝ≥0} {p : ℝ} : r < Real.toNNReal p ↔ ↑r < p := + lt_iff_lt_of_le_iff_le toNNReal_le_iff_le_coe + +theorem toNNReal_pow {x : ℝ} (hx : 0 ≤ x) (n : ℕ) : (x ^ n).toNNReal = x.toNNReal ^ n := by + rw [← coe_inj, NNReal.coe_pow, Real.coe_toNNReal _ (pow_nonneg hx _), + Real.coe_toNNReal x hx] + +theorem toNNReal_mul {p q : ℝ} (hp : 0 ≤ p) : + Real.toNNReal (p * q) = Real.toNNReal p * Real.toNNReal q := + NNReal.eq <| by simp [mul_max_of_nonneg, hp] + +end ToNNReal + +end Real + +open Real + +namespace NNReal + +section Mul + +theorem mul_eq_mul_left {a b c : ℝ≥0} (h : a ≠ 0) : a * b = a * c ↔ b = c := by + rw [mul_eq_mul_left_iff, or_iff_left h] + +end Mul + +section Pow + +theorem pow_antitone_exp {a : ℝ≥0} (m n : ℕ) (mn : m ≤ n) (a1 : a ≤ 1) : a ^ n ≤ a ^ m := + pow_le_pow_of_le_one (zero_le a) a1 mn + +nonrec theorem exists_pow_lt_of_lt_one {a b : ℝ≥0} (ha : 0 < a) (hb : b < 1) : + ∃ n : ℕ, b ^ n < a := by + simpa only [← coe_pow, NNReal.coe_lt_coe] using + exists_pow_lt_of_lt_one (NNReal.coe_pos.2 ha) (NNReal.coe_lt_coe.2 hb) + +nonrec theorem exists_mem_Ico_zpow {x : ℝ≥0} {y : ℝ≥0} (hx : x ≠ 0) (hy : 1 < y) : + ∃ n : ℤ, x ∈ Set.Ico (y ^ n) (y ^ (n + 1)) := + exists_mem_Ico_zpow (α := ℝ) hx.bot_lt hy + +nonrec theorem exists_mem_Ioc_zpow {x : ℝ≥0} {y : ℝ≥0} (hx : x ≠ 0) (hy : 1 < y) : + ∃ n : ℤ, x ∈ Set.Ioc (y ^ n) (y ^ (n + 1)) := + exists_mem_Ioc_zpow (α := ℝ) hx.bot_lt hy + +end Pow + +section Sub + +/-! +### Lemmas about subtraction + +In this section we provide a few lemmas about subtraction that do not fit well into any other +typeclass. For lemmas about subtraction and addition see lemmas about `OrderedSub` in the file +`Mathlib.Algebra.Order.Sub.Basic`. See also `mul_tsub` and `tsub_mul`. +-/ + +theorem sub_def {r p : ℝ≥0} : r - p = Real.toNNReal (r - p) := + rfl + +theorem coe_sub_def {r p : ℝ≥0} : ↑(r - p) = max (r - p : ℝ) 0 := + rfl + +example : OrderedSub ℝ≥0 := by infer_instance + +end Sub + +section Inv + +@[simp] +theorem inv_le {r p : ℝ≥0} (h : r ≠ 0) : r⁻¹ ≤ p ↔ 1 ≤ r * p := by + rw [← mul_le_mul_left (pos_iff_ne_zero.2 h), mul_inv_cancel₀ h] + +theorem inv_le_of_le_mul {r p : ℝ≥0} (h : 1 ≤ r * p) : r⁻¹ ≤ p := by + by_cases r = 0 <;> simp [*, inv_le] + +@[simp] +theorem le_inv_iff_mul_le {r p : ℝ≥0} (h : p ≠ 0) : r ≤ p⁻¹ ↔ r * p ≤ 1 := by + rw [← mul_le_mul_left (pos_iff_ne_zero.2 h), mul_inv_cancel₀ h, mul_comm] + +@[simp] +theorem lt_inv_iff_mul_lt {r p : ℝ≥0} (h : p ≠ 0) : r < p⁻¹ ↔ r * p < 1 := by + rw [← mul_lt_mul_left (pos_iff_ne_zero.2 h), mul_inv_cancel₀ h, mul_comm] + +@[deprecated le_inv_mul_iff₀ (since := "2024-08-21")] +theorem mul_le_iff_le_inv {a b r : ℝ≥0} (hr : r ≠ 0) : r * a ≤ b ↔ a ≤ r⁻¹ * b := + (le_inv_mul_iff₀ (pos_iff_ne_zero.2 hr)).symm + +@[deprecated le_div_iff₀ (since := "2024-08-21")] +theorem le_div_iff_mul_le {a b r : ℝ≥0} (hr : r ≠ 0) : a ≤ b / r ↔ a * r ≤ b := + le_div_iff₀ (pos_iff_ne_zero.2 hr) + +@[deprecated div_le_iff₀ (since := "2024-08-21")] +protected lemma div_le_iff {a b r : ℝ≥0} (hr : r ≠ 0) : a / r ≤ b ↔ a ≤ b * r := + div_le_iff₀ (pos_iff_ne_zero.2 hr) + +@[deprecated div_le_iff₀' (since := "2024-08-21")] +protected lemma div_le_iff' {a b r : ℝ≥0} (hr : r ≠ 0) : a / r ≤ b ↔ a ≤ r * b := + div_le_iff₀' (pos_iff_ne_zero.2 hr) + +theorem div_le_of_le_mul {a b c : ℝ≥0} (h : a ≤ b * c) : a / c ≤ b := + if h0 : c = 0 then by simp [h0] else (div_le_iff₀ (pos_iff_ne_zero.2 h0)).2 h + +theorem div_le_of_le_mul' {a b c : ℝ≥0} (h : a ≤ b * c) : a / b ≤ c := + div_le_of_le_mul <| mul_comm b c ▸ h + +@[deprecated le_div_iff₀ (since := "2024-08-21")] +protected lemma le_div_iff {a b r : ℝ≥0} (hr : r ≠ 0) : a ≤ b / r ↔ a * r ≤ b := + le_div_iff₀ hr.bot_lt + +@[deprecated le_div_iff₀' (since := "2024-10-02")] +theorem le_div_iff' {a b r : ℝ≥0} (hr : r ≠ 0) : a ≤ b / r ↔ r * a ≤ b := le_div_iff₀' hr.bot_lt + +@[deprecated div_lt_iff₀ (since := "2024-10-02")] +theorem div_lt_iff {a b r : ℝ≥0} (hr : r ≠ 0) : a / r < b ↔ a < b * r := div_lt_iff₀ hr.bot_lt + +@[deprecated div_lt_iff₀' (since := "2024-10-02")] +theorem div_lt_iff' {a b r : ℝ≥0} (hr : r ≠ 0) : a / r < b ↔ a < r * b := div_lt_iff₀' hr.bot_lt + +@[deprecated lt_div_iff₀ (since := "2024-10-02")] +theorem lt_div_iff {a b r : ℝ≥0} (hr : r ≠ 0) : a < b / r ↔ a * r < b := lt_div_iff₀ hr.bot_lt + +@[deprecated lt_div_iff₀' (since := "2024-10-02")] +theorem lt_div_iff' {a b r : ℝ≥0} (hr : r ≠ 0) : a < b / r ↔ r * a < b := lt_div_iff₀' hr.bot_lt + +theorem mul_lt_of_lt_div {a b r : ℝ≥0} (h : a < b / r) : a * r < b := + (lt_div_iff₀ <| pos_iff_ne_zero.2 fun hr => False.elim <| by simp [hr] at h).1 h + +theorem div_le_div_left_of_le {a b c : ℝ≥0} (c0 : c ≠ 0) (cb : c ≤ b) : + a / b ≤ a / c := + div_le_div_of_nonneg_left (zero_le _) c0.bot_lt cb + +nonrec theorem div_le_div_left {a b c : ℝ≥0} (a0 : 0 < a) (b0 : 0 < b) (c0 : 0 < c) : + a / b ≤ a / c ↔ c ≤ b := + div_le_div_left a0 b0 c0 + +theorem le_of_forall_lt_one_mul_le {x y : ℝ≥0} (h : ∀ a < 1, a * x ≤ y) : x ≤ y := + le_of_forall_ge_of_dense fun a ha => by + have hx : x ≠ 0 := pos_iff_ne_zero.1 (lt_of_le_of_lt (zero_le _) ha) + have hx' : x⁻¹ ≠ 0 := by rwa [Ne, inv_eq_zero] + have : a * x⁻¹ < 1 := by rwa [← lt_inv_iff_mul_lt hx', inv_inv] + have : a * x⁻¹ * x ≤ y := h _ this + rwa [mul_assoc, inv_mul_cancel₀ hx, mul_one] at this + +nonrec theorem half_le_self (a : ℝ≥0) : a / 2 ≤ a := + half_le_self bot_le + +nonrec theorem half_lt_self {a : ℝ≥0} (h : a ≠ 0) : a / 2 < a := + half_lt_self h.bot_lt + +theorem div_lt_one_of_lt {a b : ℝ≥0} (h : a < b) : a / b < 1 := by + rwa [div_lt_iff₀ h.bot_lt, one_mul] + +theorem _root_.Real.toNNReal_inv {x : ℝ} : Real.toNNReal x⁻¹ = (Real.toNNReal x)⁻¹ := by + rcases le_total 0 x with hx | hx + · nth_rw 1 [← Real.coe_toNNReal x hx] + rw [← NNReal.coe_inv, Real.toNNReal_coe] + · rw [toNNReal_eq_zero.mpr hx, inv_zero, toNNReal_eq_zero.mpr (inv_nonpos.mpr hx)] + +theorem _root_.Real.toNNReal_div {x y : ℝ} (hx : 0 ≤ x) : + Real.toNNReal (x / y) = Real.toNNReal x / Real.toNNReal y := by + rw [div_eq_mul_inv, div_eq_mul_inv, ← Real.toNNReal_inv, ← Real.toNNReal_mul hx] + +theorem _root_.Real.toNNReal_div' {x y : ℝ} (hy : 0 ≤ y) : + Real.toNNReal (x / y) = Real.toNNReal x / Real.toNNReal y := by + rw [div_eq_inv_mul, div_eq_inv_mul, Real.toNNReal_mul (inv_nonneg.2 hy), Real.toNNReal_inv] + +theorem inv_lt_one_iff {x : ℝ≥0} (hx : x ≠ 0) : x⁻¹ < 1 ↔ 1 < x := by + rw [← one_div, div_lt_iff₀ hx.bot_lt, one_mul] + +@[deprecated zpow_pos (since := "2024-10-08")] +protected theorem zpow_pos {x : ℝ≥0} (hx : x ≠ 0) (n : ℤ) : 0 < x ^ n := zpow_pos hx.bot_lt _ + +theorem inv_lt_inv {x y : ℝ≥0} (hx : x ≠ 0) (h : x < y) : y⁻¹ < x⁻¹ := + inv_strictAnti₀ hx.bot_lt h + +end Inv + +@[simp] +theorem abs_eq (x : ℝ≥0) : |(x : ℝ)| = x := + abs_of_nonneg x.property + +section Csupr + +open Set + +variable {ι : Sort*} {f : ι → ℝ≥0} + +theorem le_toNNReal_of_coe_le {x : ℝ≥0} {y : ℝ} (h : ↑x ≤ y) : x ≤ y.toNNReal := + (le_toNNReal_iff_coe_le <| x.2.trans h).2 h + +nonrec theorem sSup_of_not_bddAbove {s : Set ℝ≥0} (hs : ¬BddAbove s) : SupSet.sSup s = 0 := by + rw [← bddAbove_coe] at hs + rw [← coe_inj, coe_sSup, NNReal.coe_zero] + exact sSup_of_not_bddAbove hs + +theorem iSup_of_not_bddAbove (hf : ¬BddAbove (range f)) : ⨆ i, f i = 0 := + sSup_of_not_bddAbove hf + +theorem iSup_empty [IsEmpty ι] (f : ι → ℝ≥0) : ⨆ i, f i = 0 := ciSup_of_empty f + +theorem iInf_empty [IsEmpty ι] (f : ι → ℝ≥0) : ⨅ i, f i = 0 := by + rw [_root_.iInf_of_isEmpty, sInf_empty] + +@[simp] lemma iSup_eq_zero (hf : BddAbove (range f)) : ⨆ i, f i = 0 ↔ ∀ i, f i = 0 := by + cases isEmpty_or_nonempty ι + · simp + · simp [← bot_eq_zero', ← le_bot_iff, ciSup_le_iff hf] + +@[simp] +theorem iInf_const_zero {α : Sort*} : ⨅ _ : α, (0 : ℝ≥0) = 0 := by + rw [← coe_inj, coe_iInf] + exact Real.iInf_const_zero + +end Csupr + +end NNReal + +namespace Set + +namespace OrdConnected + +variable {s : Set ℝ} {t : Set ℝ≥0} + +theorem preimage_coe_nnreal_real (h : s.OrdConnected) : ((↑) ⁻¹' s : Set ℝ≥0).OrdConnected := + h.preimage_mono NNReal.coe_mono + +theorem image_coe_nnreal_real (h : t.OrdConnected) : ((↑) '' t : Set ℝ).OrdConnected := + ⟨forall_mem_image.2 fun x hx => + forall_mem_image.2 fun _y hy z hz => ⟨⟨z, x.2.trans hz.1⟩, h.out hx hy hz, rfl⟩⟩ + +-- Porting note (#11215): TODO: does it generalize to a `GaloisInsertion`? +theorem image_real_toNNReal (h : s.OrdConnected) : (Real.toNNReal '' s).OrdConnected := by + refine ⟨forall_mem_image.2 fun x hx => forall_mem_image.2 fun y hy z hz => ?_⟩ + rcases le_total y 0 with hy₀ | hy₀ + · rw [mem_Icc, Real.toNNReal_of_nonpos hy₀, nonpos_iff_eq_zero] at hz + exact ⟨y, hy, (toNNReal_of_nonpos hy₀).trans hz.2.symm⟩ + · lift y to ℝ≥0 using hy₀ + rw [toNNReal_coe] at hz + exact ⟨z, h.out hx hy ⟨toNNReal_le_iff_le_coe.1 hz.1, hz.2⟩, toNNReal_coe⟩ + +theorem preimage_real_toNNReal (h : t.OrdConnected) : (Real.toNNReal ⁻¹' t).OrdConnected := + h.preimage_mono Real.toNNReal_mono + +end OrdConnected + +end Set + +namespace Real + +/-- The absolute value on `ℝ` as a map to `ℝ≥0`. -/ +-- Porting note (kmill): `pp_nodot` has no affect here +-- unless RFC lean4#1910 leads to dot notation for CoeFun +@[pp_nodot] +def nnabs : ℝ →*₀ ℝ≥0 where + toFun x := ⟨|x|, abs_nonneg x⟩ + map_zero' := by ext; simp + map_one' := by ext; simp + map_mul' x y := by ext; simp [abs_mul] + +@[norm_cast, simp] +theorem coe_nnabs (x : ℝ) : (nnabs x : ℝ) = |x| := + rfl + +@[simp] +theorem nnabs_of_nonneg {x : ℝ} (h : 0 ≤ x) : nnabs x = toNNReal x := by + ext + rw [coe_toNNReal x h, coe_nnabs, abs_of_nonneg h] + +theorem nnabs_coe (x : ℝ≥0) : nnabs x = x := by simp + +theorem coe_toNNReal_le (x : ℝ) : (toNNReal x : ℝ) ≤ |x| := + max_le (le_abs_self _) (abs_nonneg _) + +@[simp] lemma toNNReal_abs (x : ℝ) : |x|.toNNReal = nnabs x := NNReal.coe_injective <| by simp + +theorem cast_natAbs_eq_nnabs_cast (n : ℤ) : (n.natAbs : ℝ≥0) = nnabs n := by + ext + rw [NNReal.coe_natCast, Int.cast_natAbs, Real.coe_nnabs, Int.cast_abs] + +end Real + +section StrictMono + +open NNReal + +variable {Γ₀ : Type*} [LinearOrderedCommGroupWithZero Γ₀] + +/-- If `Γ₀ˣ` is nontrivial and `f : Γ₀ →*₀ ℝ≥0` is strictly monotone, then for any positive + `r : ℝ≥0`, there exists `d : Γ₀ˣ` with `f d < r`. -/ +theorem NNReal.exists_lt_of_strictMono [h : Nontrivial Γ₀ˣ] {f : Γ₀ →*₀ ℝ≥0} (hf : StrictMono f) + {r : ℝ≥0} (hr : 0 < r) : ∃ d : Γ₀ˣ, f d < r := by + obtain ⟨g, hg1⟩ := (nontrivial_iff_exists_ne (1 : Γ₀ˣ)).mp h + set u : Γ₀ˣ := if g < 1 then g else g⁻¹ with hu + have hfu : f u < 1 := by + rw [hu] + split_ifs with hu1 + · rw [← map_one f]; exact hf hu1 + · have hfg0 : f g ≠ 0 := + fun h0 ↦ (Units.ne_zero g) ((map_eq_zero f).mp h0) + have hg1' : 1 < g := lt_of_le_of_ne (not_lt.mp hu1) hg1.symm + rw [Units.val_inv_eq_inv_val, map_inv₀, inv_lt_one_iff hfg0, ← map_one f] + exact hf hg1' + obtain ⟨n, hn⟩ := exists_pow_lt_of_lt_one hr hfu + use u ^ n + rwa [Units.val_pow_eq_pow_val, map_pow] + +/-- If `Γ₀ˣ` is nontrivial and `f : Γ₀ →*₀ ℝ≥0` is strictly monotone, then for any positive + real `r`, there exists `d : Γ₀ˣ` with `f d < r`. -/ +theorem Real.exists_lt_of_strictMono [h : Nontrivial Γ₀ˣ] {f : Γ₀ →*₀ ℝ≥0} (hf : StrictMono f) + {r : ℝ} (hr : 0 < r) : ∃ d : Γ₀ˣ, (f d : ℝ) < r := by + set s : NNReal := ⟨r, le_of_lt hr⟩ + have hs : 0 < s := hr + exact NNReal.exists_lt_of_strictMono hf hs + +end StrictMono + +namespace Mathlib.Meta.Positivity + +open Lean Meta Qq Function + +private alias ⟨_, nnreal_coe_pos⟩ := coe_pos + +/-- Extension for the `positivity` tactic: cast from `ℝ≥0` to `ℝ`. -/ +@[positivity NNReal.toReal _] +def evalNNRealtoReal : PositivityExt where eval {u α} _zα _pα e := do + match u, α, e with + | 0, ~q(ℝ), ~q(NNReal.toReal $a) => + let ra ← core q(inferInstance) q(inferInstance) a + assertInstancesCommute + match ra with + | .positive pa => pure (.positive q(nnreal_coe_pos $pa)) + | _ => pure (.nonnegative q(NNReal.coe_nonneg $a)) + | _, _, _ => throwError "not NNReal.toReal" + +end Mathlib.Meta.Positivity diff --git a/Mathlib/Data/NNReal/Star.lean b/Mathlib/Data/NNReal/Star.lean index 37a7a68423a4b..c1f11b1539a4e 100644 --- a/Mathlib/Data/NNReal/Star.lean +++ b/Mathlib/Data/NNReal/Star.lean @@ -3,8 +3,8 @@ Copyright (c) 2023 Jireh Loreaux. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jireh Loreaux -/ +import Mathlib.Data.NNReal.Defs import Mathlib.Data.Real.Star -import Mathlib.Data.NNReal.Basic /-! # The non-negative real numbers are a `*`-ring, with the trivial `*`-structure diff --git a/Mathlib/MeasureTheory/Measure/AddContent.lean b/Mathlib/MeasureTheory/Measure/AddContent.lean index e07a40a78c24f..835d2ea391739 100644 --- a/Mathlib/MeasureTheory/Measure/AddContent.lean +++ b/Mathlib/MeasureTheory/Measure/AddContent.lean @@ -3,6 +3,7 @@ Copyright (c) 2024 Rémy Degenne. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne, Peter Pfaffelhuber -/ +import Mathlib.Algebra.BigOperators.Group.Finset import Mathlib.Data.ENNReal.Basic import Mathlib.MeasureTheory.SetSemiring diff --git a/Mathlib/RingTheory/Valuation/RankOne.lean b/Mathlib/RingTheory/Valuation/RankOne.lean index f052ac9fa18e6..1180c6e8cae8c 100644 --- a/Mathlib/RingTheory/Valuation/RankOne.lean +++ b/Mathlib/RingTheory/Valuation/RankOne.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 María Inés de Frutos-Fernández. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: María Inés de Frutos-Fernández -/ -import Mathlib.Data.NNReal.Basic +import Mathlib.Data.NNReal.Defs import Mathlib.RingTheory.Valuation.Basic /-! From 2edff7137352ed19fed96f2c0b9db2a88d5b4410 Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Mon, 21 Oct 2024 09:29:01 +0000 Subject: [PATCH 228/505] chore(RingTheory/DedekindDomain/FiniteAdeleRing): remove TODO (#17951) It's done! --- Mathlib/RingTheory/DedekindDomain/FiniteAdeleRing.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/RingTheory/DedekindDomain/FiniteAdeleRing.lean b/Mathlib/RingTheory/DedekindDomain/FiniteAdeleRing.lean index 6481e7153138a..2046eac05440a 100644 --- a/Mathlib/RingTheory/DedekindDomain/FiniteAdeleRing.lean +++ b/Mathlib/RingTheory/DedekindDomain/FiniteAdeleRing.lean @@ -175,7 +175,7 @@ end FiniteIntegralAdeles /-! ### The finite adèle ring of a Dedekind domain We define the finite adèle ring of `R` as the restricted product over all maximal ideals `v` of `R` of `adicCompletion` with respect to `adicCompletionIntegers`. We prove that it is a commutative -ring. TODO: show that it is a topological ring with the restricted product topology. -/ +ring. -/ namespace ProdAdicCompletions @@ -429,7 +429,7 @@ theorem submodulesRingBasis : SubmodulesRingBasis instance : TopologicalSpace (FiniteAdeleRing R K) := SubmodulesRingBasis.topology (submodulesRingBasis R K) --- the point of the above: this now works +-- the point of `submodulesRingBasis` above: this now works example : TopologicalRing (FiniteAdeleRing R K) := inferInstance end Topology From 5195b1b9fa13843cc3e29b43b149588e937a7bd9 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 21 Oct 2024 09:29:02 +0000 Subject: [PATCH 229/505] chore(Filter/Tendsto): split from `Basic` (#17959) --- Mathlib.lean | 1 + .../Derangements/Exponential.lean | 2 +- Mathlib/Order/Filter/Basic.lean | 288 +---------------- Mathlib/Order/Filter/CardinalInter.lean | 2 +- Mathlib/Order/Filter/Extr.lean | 2 +- Mathlib/Order/Filter/Germ/Basic.lean | 2 +- Mathlib/Order/Filter/Partial.lean | 2 +- Mathlib/Order/Filter/Pi.lean | 1 + Mathlib/Order/Filter/Prod.lean | 2 +- Mathlib/Order/Filter/Tendsto.lean | 304 ++++++++++++++++++ Mathlib/Topology/Compactness/Compact.lean | 2 +- 11 files changed, 314 insertions(+), 294 deletions(-) create mode 100644 Mathlib/Order/Filter/Tendsto.lean diff --git a/Mathlib.lean b/Mathlib.lean index 6da8bfa58c2fb..c6e807c4602bf 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3711,6 +3711,7 @@ import Mathlib.Order.Filter.Prod import Mathlib.Order.Filter.Ring import Mathlib.Order.Filter.SmallSets import Mathlib.Order.Filter.Subsingleton +import Mathlib.Order.Filter.Tendsto import Mathlib.Order.Filter.Ultrafilter import Mathlib.Order.Filter.ZeroAndBoundedAtFilter import Mathlib.Order.Fin.Basic diff --git a/Mathlib/Combinatorics/Derangements/Exponential.lean b/Mathlib/Combinatorics/Derangements/Exponential.lean index 9ec2eea19d87f..f992a4f1c6ffb 100644 --- a/Mathlib/Combinatorics/Derangements/Exponential.lean +++ b/Mathlib/Combinatorics/Derangements/Exponential.lean @@ -5,7 +5,7 @@ Authors: Henry Swanson, Patrick Massot -/ import Mathlib.Analysis.SpecialFunctions.Exponential import Mathlib.Combinatorics.Derangements.Finite -import Mathlib.Order.Filter.Basic +import Mathlib.Order.Filter.Tendsto /-! # Derangement exponential series diff --git a/Mathlib/Order/Filter/Basic.lean b/Mathlib/Order/Filter/Basic.lean index 8b546db30b150..47df88df7daec 100644 --- a/Mathlib/Order/Filter/Basic.lean +++ b/Mathlib/Order/Filter/Basic.lean @@ -37,17 +37,9 @@ The examples of filters appearing in the description of the two motivating ideas * `MeasureTheory.ae` : made of sets whose complement has zero measure with respect to `μ` (defined in `Mathlib/MeasureTheory/OuterMeasure/AE`) -The general notion of limit of a map with respect to filters on the source and target types -is `Filter.Tendsto`. It is defined in terms of the order and the push-forward operation. The predicate "happening eventually" is `Filter.Eventually`, and "happening often" is `Filter.Frequently`, whose definitions are immediate after `Filter` is defined (but they come rather late in this file in order to immediately relate them to the lattice structure). - -For instance, anticipating on Topology.Basic, the statement: "if a sequence `u` converges to -some `x` and `u n` belongs to a set `M` for `n` large enough then `x` is in the closure of -`M`" is formalized as: `Tendsto u atTop (𝓝 x) → (∀ᶠ n in atTop, u n ∈ M) → x ∈ closure M`, -which is a special case of `mem_closure_of_tendsto` from Topology.Basic. - ## Notations * `∀ᶠ x in f, p x` : `f.Eventually p`; @@ -2312,263 +2304,6 @@ theorem principal_bind {s : Set α} {f : α → Filter β} : bind (𝓟 s) f = end Bind -/-! ### Limits -/ - -theorem tendsto_def {f : α → β} {l₁ : Filter α} {l₂ : Filter β} : - Tendsto f l₁ l₂ ↔ ∀ s ∈ l₂, f ⁻¹' s ∈ l₁ := - Iff.rfl - -theorem tendsto_iff_eventually {f : α → β} {l₁ : Filter α} {l₂ : Filter β} : - Tendsto f l₁ l₂ ↔ ∀ ⦃p : β → Prop⦄, (∀ᶠ y in l₂, p y) → ∀ᶠ x in l₁, p (f x) := - Iff.rfl - -theorem tendsto_iff_forall_eventually_mem {f : α → β} {l₁ : Filter α} {l₂ : Filter β} : - Tendsto f l₁ l₂ ↔ ∀ s ∈ l₂, ∀ᶠ x in l₁, f x ∈ s := - Iff.rfl - -lemma Tendsto.eventually_mem {f : α → β} {l₁ : Filter α} {l₂ : Filter β} {s : Set β} - (hf : Tendsto f l₁ l₂) (h : s ∈ l₂) : ∀ᶠ x in l₁, f x ∈ s := - hf h - -theorem Tendsto.eventually {f : α → β} {l₁ : Filter α} {l₂ : Filter β} {p : β → Prop} - (hf : Tendsto f l₁ l₂) (h : ∀ᶠ y in l₂, p y) : ∀ᶠ x in l₁, p (f x) := - hf h - -theorem not_tendsto_iff_exists_frequently_nmem {f : α → β} {l₁ : Filter α} {l₂ : Filter β} : - ¬Tendsto f l₁ l₂ ↔ ∃ s ∈ l₂, ∃ᶠ x in l₁, f x ∉ s := by - simp only [tendsto_iff_forall_eventually_mem, not_forall, exists_prop, not_eventually] - -theorem Tendsto.frequently {f : α → β} {l₁ : Filter α} {l₂ : Filter β} {p : β → Prop} - (hf : Tendsto f l₁ l₂) (h : ∃ᶠ x in l₁, p (f x)) : ∃ᶠ y in l₂, p y := - mt hf.eventually h - -theorem Tendsto.frequently_map {l₁ : Filter α} {l₂ : Filter β} {p : α → Prop} {q : β → Prop} - (f : α → β) (c : Filter.Tendsto f l₁ l₂) (w : ∀ x, p x → q (f x)) (h : ∃ᶠ x in l₁, p x) : - ∃ᶠ y in l₂, q y := - c.frequently (h.mono w) - -@[simp] -theorem tendsto_bot {f : α → β} {l : Filter β} : Tendsto f ⊥ l := by simp [Tendsto] - -theorem Tendsto.of_neBot_imp {f : α → β} {la : Filter α} {lb : Filter β} - (h : NeBot la → Tendsto f la lb) : Tendsto f la lb := by - rcases eq_or_neBot la with rfl | hla - · exact tendsto_bot - · exact h hla - -@[simp] theorem tendsto_top {f : α → β} {l : Filter α} : Tendsto f l ⊤ := le_top - -theorem le_map_of_right_inverse {mab : α → β} {mba : β → α} {f : Filter α} {g : Filter β} - (h₁ : mab ∘ mba =ᶠ[g] id) (h₂ : Tendsto mba g f) : g ≤ map mab f := by - rw [← @map_id _ g, ← map_congr h₁, ← map_map] - exact map_mono h₂ - -theorem tendsto_of_isEmpty [IsEmpty α] {f : α → β} {la : Filter α} {lb : Filter β} : - Tendsto f la lb := by simp only [filter_eq_bot_of_isEmpty la, tendsto_bot] - -theorem eventuallyEq_of_left_inv_of_right_inv {f : α → β} {g₁ g₂ : β → α} {fa : Filter α} - {fb : Filter β} (hleft : ∀ᶠ x in fa, g₁ (f x) = x) (hright : ∀ᶠ y in fb, f (g₂ y) = y) - (htendsto : Tendsto g₂ fb fa) : g₁ =ᶠ[fb] g₂ := - (htendsto.eventually hleft).mp <| hright.mono fun _ hr hl => (congr_arg g₁ hr.symm).trans hl - -theorem tendsto_iff_comap {f : α → β} {l₁ : Filter α} {l₂ : Filter β} : - Tendsto f l₁ l₂ ↔ l₁ ≤ l₂.comap f := - map_le_iff_le_comap - -alias ⟨Tendsto.le_comap, _⟩ := tendsto_iff_comap - -protected theorem Tendsto.disjoint {f : α → β} {la₁ la₂ : Filter α} {lb₁ lb₂ : Filter β} - (h₁ : Tendsto f la₁ lb₁) (hd : Disjoint lb₁ lb₂) (h₂ : Tendsto f la₂ lb₂) : Disjoint la₁ la₂ := - (disjoint_comap hd).mono h₁.le_comap h₂.le_comap - -theorem tendsto_congr' {f₁ f₂ : α → β} {l₁ : Filter α} {l₂ : Filter β} (hl : f₁ =ᶠ[l₁] f₂) : - Tendsto f₁ l₁ l₂ ↔ Tendsto f₂ l₁ l₂ := by rw [Tendsto, Tendsto, map_congr hl] - -theorem Tendsto.congr' {f₁ f₂ : α → β} {l₁ : Filter α} {l₂ : Filter β} (hl : f₁ =ᶠ[l₁] f₂) - (h : Tendsto f₁ l₁ l₂) : Tendsto f₂ l₁ l₂ := - (tendsto_congr' hl).1 h - -theorem tendsto_congr {f₁ f₂ : α → β} {l₁ : Filter α} {l₂ : Filter β} (h : ∀ x, f₁ x = f₂ x) : - Tendsto f₁ l₁ l₂ ↔ Tendsto f₂ l₁ l₂ := - tendsto_congr' (univ_mem' h) - -theorem Tendsto.congr {f₁ f₂ : α → β} {l₁ : Filter α} {l₂ : Filter β} (h : ∀ x, f₁ x = f₂ x) : - Tendsto f₁ l₁ l₂ → Tendsto f₂ l₁ l₂ := - (tendsto_congr h).1 - -theorem tendsto_id' {x y : Filter α} : Tendsto id x y ↔ x ≤ y := - Iff.rfl - -theorem tendsto_id {x : Filter α} : Tendsto id x x := - le_refl x - -theorem Tendsto.comp {f : α → β} {g : β → γ} {x : Filter α} {y : Filter β} {z : Filter γ} - (hg : Tendsto g y z) (hf : Tendsto f x y) : Tendsto (g ∘ f) x z := fun _ hs => hf (hg hs) - -protected theorem Tendsto.iterate {f : α → α} {l : Filter α} (h : Tendsto f l l) : - ∀ n, Tendsto (f^[n]) l l - | 0 => tendsto_id - | (n + 1) => (h.iterate n).comp h - -theorem Tendsto.mono_left {f : α → β} {x y : Filter α} {z : Filter β} (hx : Tendsto f x z) - (h : y ≤ x) : Tendsto f y z := - (map_mono h).trans hx - -theorem Tendsto.mono_right {f : α → β} {x : Filter α} {y z : Filter β} (hy : Tendsto f x y) - (hz : y ≤ z) : Tendsto f x z := - le_trans hy hz - -theorem Tendsto.neBot {f : α → β} {x : Filter α} {y : Filter β} (h : Tendsto f x y) [hx : NeBot x] : - NeBot y := - (hx.map _).mono h - -theorem tendsto_map {f : α → β} {x : Filter α} : Tendsto f x (map f x) := - le_refl (map f x) - -@[simp] -theorem tendsto_map'_iff {f : β → γ} {g : α → β} {x : Filter α} {y : Filter γ} : - Tendsto f (map g x) y ↔ Tendsto (f ∘ g) x y := by - rw [Tendsto, Tendsto, map_map] - -alias ⟨_, tendsto_map'⟩ := tendsto_map'_iff - -theorem tendsto_comap {f : α → β} {x : Filter β} : Tendsto f (comap f x) x := - map_comap_le - -@[simp] -theorem tendsto_comap_iff {f : α → β} {g : β → γ} {a : Filter α} {c : Filter γ} : - Tendsto f a (c.comap g) ↔ Tendsto (g ∘ f) a c := - ⟨fun h => tendsto_comap.comp h, fun h => map_le_iff_le_comap.mp <| by rwa [map_map]⟩ - -theorem tendsto_comap'_iff {m : α → β} {f : Filter α} {g : Filter β} {i : γ → α} (h : range i ∈ f) : - Tendsto (m ∘ i) (comap i f) g ↔ Tendsto m f g := by - rw [Tendsto, ← map_compose] - simp only [(· ∘ ·), map_comap_of_mem h, Tendsto] - -theorem Tendsto.of_tendsto_comp {f : α → β} {g : β → γ} {a : Filter α} {b : Filter β} {c : Filter γ} - (hfg : Tendsto (g ∘ f) a c) (hg : comap g c ≤ b) : Tendsto f a b := by - rw [tendsto_iff_comap] at hfg ⊢ - calc - a ≤ comap (g ∘ f) c := hfg - _ ≤ comap f b := by simpa [comap_comap] using comap_mono hg - -theorem comap_eq_of_inverse {f : Filter α} {g : Filter β} {φ : α → β} (ψ : β → α) (eq : ψ ∘ φ = id) - (hφ : Tendsto φ f g) (hψ : Tendsto ψ g f) : comap φ g = f := by - refine ((comap_mono <| map_le_iff_le_comap.1 hψ).trans ?_).antisymm (map_le_iff_le_comap.1 hφ) - rw [comap_comap, eq, comap_id] - -theorem map_eq_of_inverse {f : Filter α} {g : Filter β} {φ : α → β} (ψ : β → α) (eq : φ ∘ ψ = id) - (hφ : Tendsto φ f g) (hψ : Tendsto ψ g f) : map φ f = g := by - refine le_antisymm hφ (le_trans ?_ (map_mono hψ)) - rw [map_map, eq, map_id] - -theorem tendsto_inf {f : α → β} {x : Filter α} {y₁ y₂ : Filter β} : - Tendsto f x (y₁ ⊓ y₂) ↔ Tendsto f x y₁ ∧ Tendsto f x y₂ := by - simp only [Tendsto, le_inf_iff] - -theorem tendsto_inf_left {f : α → β} {x₁ x₂ : Filter α} {y : Filter β} (h : Tendsto f x₁ y) : - Tendsto f (x₁ ⊓ x₂) y := - le_trans (map_mono inf_le_left) h - -theorem tendsto_inf_right {f : α → β} {x₁ x₂ : Filter α} {y : Filter β} (h : Tendsto f x₂ y) : - Tendsto f (x₁ ⊓ x₂) y := - le_trans (map_mono inf_le_right) h - -theorem Tendsto.inf {f : α → β} {x₁ x₂ : Filter α} {y₁ y₂ : Filter β} (h₁ : Tendsto f x₁ y₁) - (h₂ : Tendsto f x₂ y₂) : Tendsto f (x₁ ⊓ x₂) (y₁ ⊓ y₂) := - tendsto_inf.2 ⟨tendsto_inf_left h₁, tendsto_inf_right h₂⟩ - -@[simp] -theorem tendsto_iInf {f : α → β} {x : Filter α} {y : ι → Filter β} : - Tendsto f x (⨅ i, y i) ↔ ∀ i, Tendsto f x (y i) := by - simp only [Tendsto, le_iInf_iff] - -theorem tendsto_iInf' {f : α → β} {x : ι → Filter α} {y : Filter β} (i : ι) - (hi : Tendsto f (x i) y) : Tendsto f (⨅ i, x i) y := - hi.mono_left <| iInf_le _ _ - -theorem tendsto_iInf_iInf {f : α → β} {x : ι → Filter α} {y : ι → Filter β} - (h : ∀ i, Tendsto f (x i) (y i)) : Tendsto f (iInf x) (iInf y) := - tendsto_iInf.2 fun i => tendsto_iInf' i (h i) - -@[simp] -theorem tendsto_sup {f : α → β} {x₁ x₂ : Filter α} {y : Filter β} : - Tendsto f (x₁ ⊔ x₂) y ↔ Tendsto f x₁ y ∧ Tendsto f x₂ y := by - simp only [Tendsto, map_sup, sup_le_iff] - -theorem Tendsto.sup {f : α → β} {x₁ x₂ : Filter α} {y : Filter β} : - Tendsto f x₁ y → Tendsto f x₂ y → Tendsto f (x₁ ⊔ x₂) y := fun h₁ h₂ => tendsto_sup.mpr ⟨h₁, h₂⟩ - -theorem Tendsto.sup_sup {f : α → β} {x₁ x₂ : Filter α} {y₁ y₂ : Filter β} - (h₁ : Tendsto f x₁ y₁) (h₂ : Tendsto f x₂ y₂) : Tendsto f (x₁ ⊔ x₂) (y₁ ⊔ y₂) := - tendsto_sup.mpr ⟨h₁.mono_right le_sup_left, h₂.mono_right le_sup_right⟩ - -@[simp] -theorem tendsto_iSup {f : α → β} {x : ι → Filter α} {y : Filter β} : - Tendsto f (⨆ i, x i) y ↔ ∀ i, Tendsto f (x i) y := by simp only [Tendsto, map_iSup, iSup_le_iff] - -theorem tendsto_iSup_iSup {f : α → β} {x : ι → Filter α} {y : ι → Filter β} - (h : ∀ i, Tendsto f (x i) (y i)) : Tendsto f (iSup x) (iSup y) := - tendsto_iSup.2 fun i => (h i).mono_right <| le_iSup _ _ - -@[simp] theorem tendsto_principal {f : α → β} {l : Filter α} {s : Set β} : - Tendsto f l (𝓟 s) ↔ ∀ᶠ a in l, f a ∈ s := by - simp only [Tendsto, le_principal_iff, mem_map', Filter.Eventually] - --- Porting note: was a `simp` lemma -theorem tendsto_principal_principal {f : α → β} {s : Set α} {t : Set β} : - Tendsto f (𝓟 s) (𝓟 t) ↔ ∀ a ∈ s, f a ∈ t := by - simp only [tendsto_principal, eventually_principal] - -@[simp] theorem tendsto_pure {f : α → β} {a : Filter α} {b : β} : - Tendsto f a (pure b) ↔ ∀ᶠ x in a, f x = b := by - simp only [Tendsto, le_pure_iff, mem_map', mem_singleton_iff, Filter.Eventually] - -theorem tendsto_pure_pure (f : α → β) (a : α) : Tendsto f (pure a) (pure (f a)) := - tendsto_pure.2 rfl - -theorem tendsto_const_pure {a : Filter α} {b : β} : Tendsto (fun _ => b) a (pure b) := - tendsto_pure.2 <| univ_mem' fun _ => rfl - -theorem pure_le_iff {a : α} {l : Filter α} : pure a ≤ l ↔ ∀ s ∈ l, a ∈ s := - Iff.rfl - -theorem tendsto_pure_left {f : α → β} {a : α} {l : Filter β} : - Tendsto f (pure a) l ↔ ∀ s ∈ l, f a ∈ s := - Iff.rfl - -@[simp] -theorem map_inf_principal_preimage {f : α → β} {s : Set β} {l : Filter α} : - map f (l ⊓ 𝓟 (f ⁻¹' s)) = map f l ⊓ 𝓟 s := - Filter.ext fun t => by simp only [mem_map', mem_inf_principal, mem_setOf_eq, mem_preimage] - -/-- If two filters are disjoint, then a function cannot tend to both of them along a non-trivial -filter. -/ -theorem Tendsto.not_tendsto {f : α → β} {a : Filter α} {b₁ b₂ : Filter β} (hf : Tendsto f a b₁) - [NeBot a] (hb : Disjoint b₁ b₂) : ¬Tendsto f a b₂ := fun hf' => - (tendsto_inf.2 ⟨hf, hf'⟩).neBot.ne hb.eq_bot - -protected theorem Tendsto.if {l₁ : Filter α} {l₂ : Filter β} {f g : α → β} {p : α → Prop} - [∀ x, Decidable (p x)] (h₀ : Tendsto f (l₁ ⊓ 𝓟 { x | p x }) l₂) - (h₁ : Tendsto g (l₁ ⊓ 𝓟 { x | ¬p x }) l₂) : - Tendsto (fun x => if p x then f x else g x) l₁ l₂ := by - simp only [tendsto_def, mem_inf_principal] at * - intro s hs - filter_upwards [h₀ s hs, h₁ s hs] with x hp₀ hp₁ - rw [mem_preimage] - split_ifs with h - exacts [hp₀ h, hp₁ h] - -protected theorem Tendsto.if' {α β : Type*} {l₁ : Filter α} {l₂ : Filter β} {f g : α → β} - {p : α → Prop} [DecidablePred p] (hf : Tendsto f l₁ l₂) (hg : Tendsto g l₁ l₂) : - Tendsto (fun a => if p a then f a else g a) l₁ l₂ := - (tendsto_inf_left hf).if (tendsto_inf_left hg) - -protected theorem Tendsto.piecewise {l₁ : Filter α} {l₂ : Filter β} {f g : α → β} {s : Set α} - [∀ x, Decidable (x ∈ s)] (h₀ : Tendsto f (l₁ ⊓ 𝓟 s) l₂) (h₁ : Tendsto g (l₁ ⊓ 𝓟 sᶜ) l₂) : - Tendsto (piecewise s f g) l₁ l₂ := - Tendsto.if h₀ h₁ - end Filter open Filter @@ -2583,29 +2318,8 @@ theorem Set.EqOn.eventuallyEq_of_mem {α β} {s : Set α} {l : Filter α} {f g : theorem HasSubset.Subset.eventuallyLE {α} {l : Filter α} {s t : Set α} (h : s ⊆ t) : s ≤ᶠ[l] t := Filter.Eventually.of_forall h -theorem Set.MapsTo.tendsto {α β} {s : Set α} {t : Set β} {f : α → β} (h : MapsTo f s t) : - Filter.Tendsto f (𝓟 s) (𝓟 t) := - Filter.tendsto_principal_principal.2 h - -theorem Filter.EventuallyEq.comp_tendsto {α β γ : Type*} {l : Filter α} {f : α → β} {f' : α → β} - (H : f =ᶠ[l] f') {g : γ → α} {lc : Filter γ} (hg : Tendsto g lc l) : - f ∘ g =ᶠ[lc] f' ∘ g := - hg.eventually H - variable {α β : Type*} {F : Filter α} {G : Filter β} -theorem Filter.map_mapsTo_Iic_iff_tendsto {m : α → β} : - MapsTo (map m) (Iic F) (Iic G) ↔ Tendsto m F G := - ⟨fun hm ↦ hm right_mem_Iic, fun hm _ ↦ hm.mono_left⟩ - -alias ⟨_, Filter.Tendsto.map_mapsTo_Iic⟩ := Filter.map_mapsTo_Iic_iff_tendsto - -theorem Filter.map_mapsTo_Iic_iff_mapsTo {s : Set α} {t : Set β} {m : α → β} : - MapsTo (map m) (Iic <| 𝓟 s) (Iic <| 𝓟 t) ↔ MapsTo m s t := by - rw [map_mapsTo_Iic_iff_tendsto, tendsto_principal_principal, MapsTo] - -alias ⟨_, Set.MapsTo.filter_map_Iic⟩ := Filter.map_mapsTo_Iic_iff_mapsTo - -- TODO(Anatole): unify with the global case theorem Filter.map_surjOn_Iic_iff_le_map {m : α → β} : SurjOn (map m) (Iic F) (Iic G) ↔ G ≤ map m F := by @@ -2639,4 +2353,4 @@ lemma compl_mem_comk {p : Set α → Prop} {he hmono hunion s} : end Filter -set_option linter.style.longFile 2800 +set_option linter.style.longFile 2500 diff --git a/Mathlib/Order/Filter/CardinalInter.lean b/Mathlib/Order/Filter/CardinalInter.lean index def8ba80b0a80..bfce10f3629f3 100644 --- a/Mathlib/Order/Filter/CardinalInter.lean +++ b/Mathlib/Order/Filter/CardinalInter.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 Josha Dekker. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Josha Dekker -/ -import Mathlib.Order.Filter.Basic +import Mathlib.Order.Filter.Tendsto import Mathlib.Order.Filter.CountableInter import Mathlib.SetTheory.Cardinal.Arithmetic import Mathlib.SetTheory.Cardinal.Cofinality diff --git a/Mathlib/Order/Filter/Extr.lean b/Mathlib/Order/Filter/Extr.lean index aa4829ffbd4a4..e4dbcb43983c5 100644 --- a/Mathlib/Order/Filter/Extr.lean +++ b/Mathlib/Order/Filter/Extr.lean @@ -3,7 +3,7 @@ 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.Order.Filter.Basic +import Mathlib.Order.Filter.Tendsto import Mathlib.Order.ConditionallyCompleteLattice.Basic import Mathlib.Algebra.Order.Group.Defs diff --git a/Mathlib/Order/Filter/Germ/Basic.lean b/Mathlib/Order/Filter/Germ/Basic.lean index 30542abc2fd0b..19a910e7b0fdb 100644 --- a/Mathlib/Order/Filter/Germ/Basic.lean +++ b/Mathlib/Order/Filter/Germ/Basic.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov, Abhimanyu Pallavi Sudhir -/ -import Mathlib.Order.Filter.Basic +import Mathlib.Order.Filter.Tendsto import Mathlib.Algebra.Module.Pi /-! diff --git a/Mathlib/Order/Filter/Partial.lean b/Mathlib/Order/Filter/Partial.lean index 05453a60c9e88..b09c3cd41ecad 100644 --- a/Mathlib/Order/Filter/Partial.lean +++ b/Mathlib/Order/Filter/Partial.lean @@ -3,7 +3,7 @@ Copyright (c) 2019 Jeremy Avigad. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad -/ -import Mathlib.Order.Filter.Basic +import Mathlib.Order.Filter.Tendsto import Mathlib.Data.PFun /-! diff --git a/Mathlib/Order/Filter/Pi.lean b/Mathlib/Order/Filter/Pi.lean index 50cf2341b3b4e..df2bc56ec23bb 100644 --- a/Mathlib/Order/Filter/Pi.lean +++ b/Mathlib/Order/Filter/Pi.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov, Alex Kontorovich -/ import Mathlib.Order.Filter.Bases +import Mathlib.Order.Filter.Tendsto /-! # (Co)product of a family of filters diff --git a/Mathlib/Order/Filter/Prod.lean b/Mathlib/Order/Filter/Prod.lean index b2af14e42295d..5a1d1ab10faa6 100644 --- a/Mathlib/Order/Filter/Prod.lean +++ b/Mathlib/Order/Filter/Prod.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Patrick Massot, Yury Kudryashov, Kevin H. Wilson, Heather Macbeth -/ -import Mathlib.Order.Filter.Basic +import Mathlib.Order.Filter.Tendsto /-! # Product and coproduct filters diff --git a/Mathlib/Order/Filter/Tendsto.lean b/Mathlib/Order/Filter/Tendsto.lean new file mode 100644 index 0000000000000..64e557baca965 --- /dev/null +++ b/Mathlib/Order/Filter/Tendsto.lean @@ -0,0 +1,304 @@ +/- +Copyright (c) 2017 Johannes Hölzl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johannes Hölzl, Jeremy Avigad +-/ +import Mathlib.Order.Filter.Basic + +/-! +# Convergence in terms of filters + +The general notion of limit of a map with respect to filters on the source and target types +is `Filter.Tendsto`. It is defined in terms of the order and the push-forward operation. + +For instance, anticipating on Topology.Basic, the statement: "if a sequence `u` converges to +some `x` and `u n` belongs to a set `M` for `n` large enough then `x` is in the closure of +`M`" is formalized as: `Tendsto u atTop (𝓝 x) → (∀ᶠ n in atTop, u n ∈ M) → x ∈ closure M`, +which is a special case of `mem_closure_of_tendsto` from `Topology/Basic`. +-/ + +open Set Filter + +variable {α β γ : Type*} {ι : Sort*} + +namespace Filter + +theorem tendsto_def {f : α → β} {l₁ : Filter α} {l₂ : Filter β} : + Tendsto f l₁ l₂ ↔ ∀ s ∈ l₂, f ⁻¹' s ∈ l₁ := + Iff.rfl + +theorem tendsto_iff_eventually {f : α → β} {l₁ : Filter α} {l₂ : Filter β} : + Tendsto f l₁ l₂ ↔ ∀ ⦃p : β → Prop⦄, (∀ᶠ y in l₂, p y) → ∀ᶠ x in l₁, p (f x) := + Iff.rfl + +theorem tendsto_iff_forall_eventually_mem {f : α → β} {l₁ : Filter α} {l₂ : Filter β} : + Tendsto f l₁ l₂ ↔ ∀ s ∈ l₂, ∀ᶠ x in l₁, f x ∈ s := + Iff.rfl + +lemma Tendsto.eventually_mem {f : α → β} {l₁ : Filter α} {l₂ : Filter β} {s : Set β} + (hf : Tendsto f l₁ l₂) (h : s ∈ l₂) : ∀ᶠ x in l₁, f x ∈ s := + hf h + +theorem Tendsto.eventually {f : α → β} {l₁ : Filter α} {l₂ : Filter β} {p : β → Prop} + (hf : Tendsto f l₁ l₂) (h : ∀ᶠ y in l₂, p y) : ∀ᶠ x in l₁, p (f x) := + hf h + +theorem not_tendsto_iff_exists_frequently_nmem {f : α → β} {l₁ : Filter α} {l₂ : Filter β} : + ¬Tendsto f l₁ l₂ ↔ ∃ s ∈ l₂, ∃ᶠ x in l₁, f x ∉ s := by + simp only [tendsto_iff_forall_eventually_mem, not_forall, exists_prop, not_eventually] + +theorem Tendsto.frequently {f : α → β} {l₁ : Filter α} {l₂ : Filter β} {p : β → Prop} + (hf : Tendsto f l₁ l₂) (h : ∃ᶠ x in l₁, p (f x)) : ∃ᶠ y in l₂, p y := + mt hf.eventually h + +theorem Tendsto.frequently_map {l₁ : Filter α} {l₂ : Filter β} {p : α → Prop} {q : β → Prop} + (f : α → β) (c : Filter.Tendsto f l₁ l₂) (w : ∀ x, p x → q (f x)) (h : ∃ᶠ x in l₁, p x) : + ∃ᶠ y in l₂, q y := + c.frequently (h.mono w) + +@[simp] +theorem tendsto_bot {f : α → β} {l : Filter β} : Tendsto f ⊥ l := by simp [Tendsto] + +theorem Tendsto.of_neBot_imp {f : α → β} {la : Filter α} {lb : Filter β} + (h : NeBot la → Tendsto f la lb) : Tendsto f la lb := by + rcases eq_or_neBot la with rfl | hla + · exact tendsto_bot + · exact h hla + +@[simp] theorem tendsto_top {f : α → β} {l : Filter α} : Tendsto f l ⊤ := le_top + +theorem le_map_of_right_inverse {mab : α → β} {mba : β → α} {f : Filter α} {g : Filter β} + (h₁ : mab ∘ mba =ᶠ[g] id) (h₂ : Tendsto mba g f) : g ≤ map mab f := by + rw [← @map_id _ g, ← map_congr h₁, ← map_map] + exact map_mono h₂ + +theorem tendsto_of_isEmpty [IsEmpty α] {f : α → β} {la : Filter α} {lb : Filter β} : + Tendsto f la lb := by simp only [filter_eq_bot_of_isEmpty la, tendsto_bot] + +theorem eventuallyEq_of_left_inv_of_right_inv {f : α → β} {g₁ g₂ : β → α} {fa : Filter α} + {fb : Filter β} (hleft : ∀ᶠ x in fa, g₁ (f x) = x) (hright : ∀ᶠ y in fb, f (g₂ y) = y) + (htendsto : Tendsto g₂ fb fa) : g₁ =ᶠ[fb] g₂ := + (htendsto.eventually hleft).mp <| hright.mono fun _ hr hl => (congr_arg g₁ hr.symm).trans hl + +theorem tendsto_iff_comap {f : α → β} {l₁ : Filter α} {l₂ : Filter β} : + Tendsto f l₁ l₂ ↔ l₁ ≤ l₂.comap f := + map_le_iff_le_comap + +alias ⟨Tendsto.le_comap, _⟩ := tendsto_iff_comap + +protected theorem Tendsto.disjoint {f : α → β} {la₁ la₂ : Filter α} {lb₁ lb₂ : Filter β} + (h₁ : Tendsto f la₁ lb₁) (hd : Disjoint lb₁ lb₂) (h₂ : Tendsto f la₂ lb₂) : Disjoint la₁ la₂ := + (disjoint_comap hd).mono h₁.le_comap h₂.le_comap + +theorem tendsto_congr' {f₁ f₂ : α → β} {l₁ : Filter α} {l₂ : Filter β} (hl : f₁ =ᶠ[l₁] f₂) : + Tendsto f₁ l₁ l₂ ↔ Tendsto f₂ l₁ l₂ := by rw [Tendsto, Tendsto, map_congr hl] + +theorem Tendsto.congr' {f₁ f₂ : α → β} {l₁ : Filter α} {l₂ : Filter β} (hl : f₁ =ᶠ[l₁] f₂) + (h : Tendsto f₁ l₁ l₂) : Tendsto f₂ l₁ l₂ := + (tendsto_congr' hl).1 h + +theorem tendsto_congr {f₁ f₂ : α → β} {l₁ : Filter α} {l₂ : Filter β} (h : ∀ x, f₁ x = f₂ x) : + Tendsto f₁ l₁ l₂ ↔ Tendsto f₂ l₁ l₂ := + tendsto_congr' (univ_mem' h) + +theorem Tendsto.congr {f₁ f₂ : α → β} {l₁ : Filter α} {l₂ : Filter β} (h : ∀ x, f₁ x = f₂ x) : + Tendsto f₁ l₁ l₂ → Tendsto f₂ l₁ l₂ := + (tendsto_congr h).1 + +theorem tendsto_id' {x y : Filter α} : Tendsto id x y ↔ x ≤ y := + Iff.rfl + +theorem tendsto_id {x : Filter α} : Tendsto id x x := + le_refl x + +theorem Tendsto.comp {f : α → β} {g : β → γ} {x : Filter α} {y : Filter β} {z : Filter γ} + (hg : Tendsto g y z) (hf : Tendsto f x y) : Tendsto (g ∘ f) x z := fun _ hs => hf (hg hs) + +protected theorem Tendsto.iterate {f : α → α} {l : Filter α} (h : Tendsto f l l) : + ∀ n, Tendsto (f^[n]) l l + | 0 => tendsto_id + | (n + 1) => (h.iterate n).comp h + +theorem Tendsto.mono_left {f : α → β} {x y : Filter α} {z : Filter β} (hx : Tendsto f x z) + (h : y ≤ x) : Tendsto f y z := + (map_mono h).trans hx + +theorem Tendsto.mono_right {f : α → β} {x : Filter α} {y z : Filter β} (hy : Tendsto f x y) + (hz : y ≤ z) : Tendsto f x z := + le_trans hy hz + +theorem Tendsto.neBot {f : α → β} {x : Filter α} {y : Filter β} (h : Tendsto f x y) [hx : NeBot x] : + NeBot y := + (hx.map _).mono h + +theorem tendsto_map {f : α → β} {x : Filter α} : Tendsto f x (map f x) := + le_refl (map f x) + +@[simp] +theorem tendsto_map'_iff {f : β → γ} {g : α → β} {x : Filter α} {y : Filter γ} : + Tendsto f (map g x) y ↔ Tendsto (f ∘ g) x y := by + rw [Tendsto, Tendsto, map_map] + +alias ⟨_, tendsto_map'⟩ := tendsto_map'_iff + +theorem tendsto_comap {f : α → β} {x : Filter β} : Tendsto f (comap f x) x := + map_comap_le + +@[simp] +theorem tendsto_comap_iff {f : α → β} {g : β → γ} {a : Filter α} {c : Filter γ} : + Tendsto f a (c.comap g) ↔ Tendsto (g ∘ f) a c := + ⟨fun h => tendsto_comap.comp h, fun h => map_le_iff_le_comap.mp <| by rwa [map_map]⟩ + +theorem tendsto_comap'_iff {m : α → β} {f : Filter α} {g : Filter β} {i : γ → α} (h : range i ∈ f) : + Tendsto (m ∘ i) (comap i f) g ↔ Tendsto m f g := by + rw [Tendsto, ← map_compose] + simp only [(· ∘ ·), map_comap_of_mem h, Tendsto] + +theorem Tendsto.of_tendsto_comp {f : α → β} {g : β → γ} {a : Filter α} {b : Filter β} {c : Filter γ} + (hfg : Tendsto (g ∘ f) a c) (hg : comap g c ≤ b) : Tendsto f a b := by + rw [tendsto_iff_comap] at hfg ⊢ + calc + a ≤ comap (g ∘ f) c := hfg + _ ≤ comap f b := by simpa [comap_comap] using comap_mono hg + +theorem comap_eq_of_inverse {f : Filter α} {g : Filter β} {φ : α → β} (ψ : β → α) (eq : ψ ∘ φ = id) + (hφ : Tendsto φ f g) (hψ : Tendsto ψ g f) : comap φ g = f := by + refine ((comap_mono <| map_le_iff_le_comap.1 hψ).trans ?_).antisymm (map_le_iff_le_comap.1 hφ) + rw [comap_comap, eq, comap_id] + +theorem map_eq_of_inverse {f : Filter α} {g : Filter β} {φ : α → β} (ψ : β → α) (eq : φ ∘ ψ = id) + (hφ : Tendsto φ f g) (hψ : Tendsto ψ g f) : map φ f = g := by + refine le_antisymm hφ (le_trans ?_ (map_mono hψ)) + rw [map_map, eq, map_id] + +theorem tendsto_inf {f : α → β} {x : Filter α} {y₁ y₂ : Filter β} : + Tendsto f x (y₁ ⊓ y₂) ↔ Tendsto f x y₁ ∧ Tendsto f x y₂ := by + simp only [Tendsto, le_inf_iff] + +theorem tendsto_inf_left {f : α → β} {x₁ x₂ : Filter α} {y : Filter β} (h : Tendsto f x₁ y) : + Tendsto f (x₁ ⊓ x₂) y := + le_trans (map_mono inf_le_left) h + +theorem tendsto_inf_right {f : α → β} {x₁ x₂ : Filter α} {y : Filter β} (h : Tendsto f x₂ y) : + Tendsto f (x₁ ⊓ x₂) y := + le_trans (map_mono inf_le_right) h + +theorem Tendsto.inf {f : α → β} {x₁ x₂ : Filter α} {y₁ y₂ : Filter β} (h₁ : Tendsto f x₁ y₁) + (h₂ : Tendsto f x₂ y₂) : Tendsto f (x₁ ⊓ x₂) (y₁ ⊓ y₂) := + tendsto_inf.2 ⟨tendsto_inf_left h₁, tendsto_inf_right h₂⟩ + +@[simp] +theorem tendsto_iInf {f : α → β} {x : Filter α} {y : ι → Filter β} : + Tendsto f x (⨅ i, y i) ↔ ∀ i, Tendsto f x (y i) := by + simp only [Tendsto, le_iInf_iff] + +theorem tendsto_iInf' {f : α → β} {x : ι → Filter α} {y : Filter β} (i : ι) + (hi : Tendsto f (x i) y) : Tendsto f (⨅ i, x i) y := + hi.mono_left <| iInf_le _ _ + +theorem tendsto_iInf_iInf {f : α → β} {x : ι → Filter α} {y : ι → Filter β} + (h : ∀ i, Tendsto f (x i) (y i)) : Tendsto f (iInf x) (iInf y) := + tendsto_iInf.2 fun i => tendsto_iInf' i (h i) + +@[simp] +theorem tendsto_sup {f : α → β} {x₁ x₂ : Filter α} {y : Filter β} : + Tendsto f (x₁ ⊔ x₂) y ↔ Tendsto f x₁ y ∧ Tendsto f x₂ y := by + simp only [Tendsto, map_sup, sup_le_iff] + +theorem Tendsto.sup {f : α → β} {x₁ x₂ : Filter α} {y : Filter β} : + Tendsto f x₁ y → Tendsto f x₂ y → Tendsto f (x₁ ⊔ x₂) y := fun h₁ h₂ => tendsto_sup.mpr ⟨h₁, h₂⟩ + +theorem Tendsto.sup_sup {f : α → β} {x₁ x₂ : Filter α} {y₁ y₂ : Filter β} + (h₁ : Tendsto f x₁ y₁) (h₂ : Tendsto f x₂ y₂) : Tendsto f (x₁ ⊔ x₂) (y₁ ⊔ y₂) := + tendsto_sup.mpr ⟨h₁.mono_right le_sup_left, h₂.mono_right le_sup_right⟩ + +@[simp] +theorem tendsto_iSup {f : α → β} {x : ι → Filter α} {y : Filter β} : + Tendsto f (⨆ i, x i) y ↔ ∀ i, Tendsto f (x i) y := by simp only [Tendsto, map_iSup, iSup_le_iff] + +theorem tendsto_iSup_iSup {f : α → β} {x : ι → Filter α} {y : ι → Filter β} + (h : ∀ i, Tendsto f (x i) (y i)) : Tendsto f (iSup x) (iSup y) := + tendsto_iSup.2 fun i => (h i).mono_right <| le_iSup _ _ + +@[simp] theorem tendsto_principal {f : α → β} {l : Filter α} {s : Set β} : + Tendsto f l (𝓟 s) ↔ ∀ᶠ a in l, f a ∈ s := by + simp only [Tendsto, le_principal_iff, mem_map', Filter.Eventually] + +-- Porting note: was a `simp` lemma +theorem tendsto_principal_principal {f : α → β} {s : Set α} {t : Set β} : + Tendsto f (𝓟 s) (𝓟 t) ↔ ∀ a ∈ s, f a ∈ t := by + simp only [tendsto_principal, eventually_principal] + +@[simp] theorem tendsto_pure {f : α → β} {a : Filter α} {b : β} : + Tendsto f a (pure b) ↔ ∀ᶠ x in a, f x = b := by + simp only [Tendsto, le_pure_iff, mem_map', mem_singleton_iff, Filter.Eventually] + +theorem tendsto_pure_pure (f : α → β) (a : α) : Tendsto f (pure a) (pure (f a)) := + tendsto_pure.2 rfl + +theorem tendsto_const_pure {a : Filter α} {b : β} : Tendsto (fun _ => b) a (pure b) := + tendsto_pure.2 <| univ_mem' fun _ => rfl + +theorem pure_le_iff {a : α} {l : Filter α} : pure a ≤ l ↔ ∀ s ∈ l, a ∈ s := + Iff.rfl + +theorem tendsto_pure_left {f : α → β} {a : α} {l : Filter β} : + Tendsto f (pure a) l ↔ ∀ s ∈ l, f a ∈ s := + Iff.rfl + +@[simp] +theorem map_inf_principal_preimage {f : α → β} {s : Set β} {l : Filter α} : + map f (l ⊓ 𝓟 (f ⁻¹' s)) = map f l ⊓ 𝓟 s := + Filter.ext fun t => by simp only [mem_map', mem_inf_principal, mem_setOf_eq, mem_preimage] + +/-- If two filters are disjoint, then a function cannot tend to both of them along a non-trivial +filter. -/ +theorem Tendsto.not_tendsto {f : α → β} {a : Filter α} {b₁ b₂ : Filter β} (hf : Tendsto f a b₁) + [NeBot a] (hb : Disjoint b₁ b₂) : ¬Tendsto f a b₂ := fun hf' => + (tendsto_inf.2 ⟨hf, hf'⟩).neBot.ne hb.eq_bot + +protected theorem Tendsto.if {l₁ : Filter α} {l₂ : Filter β} {f g : α → β} {p : α → Prop} + [∀ x, Decidable (p x)] (h₀ : Tendsto f (l₁ ⊓ 𝓟 { x | p x }) l₂) + (h₁ : Tendsto g (l₁ ⊓ 𝓟 { x | ¬p x }) l₂) : + Tendsto (fun x => if p x then f x else g x) l₁ l₂ := by + simp only [tendsto_def, mem_inf_principal] at * + intro s hs + filter_upwards [h₀ s hs, h₁ s hs] with x hp₀ hp₁ + rw [mem_preimage] + split_ifs with h + exacts [hp₀ h, hp₁ h] + +protected theorem Tendsto.if' {α β : Type*} {l₁ : Filter α} {l₂ : Filter β} {f g : α → β} + {p : α → Prop} [DecidablePred p] (hf : Tendsto f l₁ l₂) (hg : Tendsto g l₁ l₂) : + Tendsto (fun a => if p a then f a else g a) l₁ l₂ := + (tendsto_inf_left hf).if (tendsto_inf_left hg) + +protected theorem Tendsto.piecewise {l₁ : Filter α} {l₂ : Filter β} {f g : α → β} {s : Set α} + [∀ x, Decidable (x ∈ s)] (h₀ : Tendsto f (l₁ ⊓ 𝓟 s) l₂) (h₁ : Tendsto g (l₁ ⊓ 𝓟 sᶜ) l₂) : + Tendsto (piecewise s f g) l₁ l₂ := + Tendsto.if h₀ h₁ + +end Filter + +theorem Set.MapsTo.tendsto {s : Set α} {t : Set β} {f : α → β} (h : MapsTo f s t) : + Filter.Tendsto f (𝓟 s) (𝓟 t) := + Filter.tendsto_principal_principal.2 h + +theorem Filter.EventuallyEq.comp_tendsto {l : Filter α} {f : α → β} {f' : α → β} + (H : f =ᶠ[l] f') {g : γ → α} {lc : Filter γ} (hg : Tendsto g lc l) : + f ∘ g =ᶠ[lc] f' ∘ g := + hg.eventually H + +variable {F : Filter α} {G : Filter β} + +theorem Filter.map_mapsTo_Iic_iff_tendsto {m : α → β} : + MapsTo (map m) (Iic F) (Iic G) ↔ Tendsto m F G := + ⟨fun hm ↦ hm right_mem_Iic, fun hm _ ↦ hm.mono_left⟩ + +alias ⟨_, Filter.Tendsto.map_mapsTo_Iic⟩ := Filter.map_mapsTo_Iic_iff_tendsto + +theorem Filter.map_mapsTo_Iic_iff_mapsTo {s : Set α} {t : Set β} {m : α → β} : + MapsTo (map m) (Iic <| 𝓟 s) (Iic <| 𝓟 t) ↔ MapsTo m s t := by + rw [map_mapsTo_Iic_iff_tendsto, tendsto_principal_principal, MapsTo] + +alias ⟨_, Set.MapsTo.filter_map_Iic⟩ := Filter.map_mapsTo_Iic_iff_mapsTo diff --git a/Mathlib/Topology/Compactness/Compact.lean b/Mathlib/Topology/Compactness/Compact.lean index ac4cbe1c1f71a..306e7735c660f 100644 --- a/Mathlib/Topology/Compactness/Compact.lean +++ b/Mathlib/Topology/Compactness/Compact.lean @@ -3,7 +3,7 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro, Yury Kudryashov -/ -import Mathlib.Order.Filter.Basic +import Mathlib.Order.Filter.Tendsto import Mathlib.Topology.Bases import Mathlib.Data.Set.Accumulate import Mathlib.Topology.Bornology.Basic From 336b74698c77f771b9c02a8edfdad574ec729752 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Mon, 21 Oct 2024 09:29:04 +0000 Subject: [PATCH 230/505] feat(Topology/Inseparable): add `alias Specializes.of_eq := specializes_of_eq` (#17988) To enable dot notation. --- Mathlib/Topology/Inseparable.lean | 2 ++ 1 file changed, 2 insertions(+) diff --git a/Mathlib/Topology/Inseparable.lean b/Mathlib/Topology/Inseparable.lean index 85b288a5061a8..919eeb7228337 100644 --- a/Mathlib/Topology/Inseparable.lean +++ b/Mathlib/Topology/Inseparable.lean @@ -132,6 +132,8 @@ theorem Specializes.trans : x ⤳ y → y ⤳ z → x ⤳ z := theorem specializes_of_eq (e : x = y) : x ⤳ y := e ▸ specializes_refl x +alias Specializes.of_eq := specializes_of_eq + theorem specializes_of_nhdsWithin (h₁ : 𝓝[s] x ≤ 𝓝[s] y) (h₂ : x ∈ s) : x ⤳ y := specializes_iff_pure.2 <| calc From 355721b6f6d31cfbbf74c79daaa5b52de44594e6 Mon Sep 17 00:00:00 2001 From: "Tristan F.-R." Date: Mon, 21 Oct 2024 09:29:05 +0000 Subject: [PATCH 231/505] docs(Computability/EpsilonNFA): typo (#17992) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Fixes a typo in `εNFA` docs (inputing -> inputting) Co-authored-by: Tristan F. --- Mathlib/Computability/EpsilonNFA.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Computability/EpsilonNFA.lean b/Mathlib/Computability/EpsilonNFA.lean index 15f0b154b182a..2afdb82ebebaa 100644 --- a/Mathlib/Computability/EpsilonNFA.lean +++ b/Mathlib/Computability/EpsilonNFA.lean @@ -28,7 +28,7 @@ universe u v /-- An `εNFA` is a set of states (`σ`), a transition function from state to state labelled by the alphabet (`step`), a starting state (`start`) and a set of acceptance states (`accept`). Note the transition function sends a state to a `Set` of states and can make ε-transitions by - inputing `none`. + inputting `none`. Since this definition allows for Automata with infinite states, a `Fintype` instance must be supplied for true `εNFA`'s. -/ structure εNFA (α : Type u) (σ : Type v) where From 59834f406e6ac3402a2efea724a777a5720ea319 Mon Sep 17 00:00:00 2001 From: damiano Date: Mon, 21 Oct 2024 09:56:17 +0000 Subject: [PATCH 232/505] chore: removed unused variables (#17975) Part of/working towards #17715 --- Mathlib/Data/Matrix/Hadamard.lean | 5 ++--- Mathlib/Data/Matrix/Reflection.lean | 2 +- Mathlib/Data/Matroid/Closure.lean | 2 +- Mathlib/Data/Matroid/IndepAxioms.lean | 2 +- Mathlib/Data/Matroid/Map.lean | 2 +- Mathlib/Data/Matroid/Restrict.lean | 2 +- Mathlib/Data/NNRat/BigOperators.lean | 2 +- Mathlib/Data/Nat/Factorial/SuperFactorial.lean | 2 -- Mathlib/Data/Nat/Fib/Zeckendorf.lean | 2 +- Mathlib/Data/Sum/Interval.lean | 2 +- Mathlib/Deprecated/Subgroup.lean | 2 +- Mathlib/Deprecated/Subring.lean | 2 -- Mathlib/FieldTheory/Finite/Polynomial.lean | 6 ------ Mathlib/FieldTheory/IsAlgClosed/Classification.lean | 4 +--- Mathlib/FieldTheory/IsPerfectClosure.lean | 2 +- Mathlib/Geometry/Euclidean/Angle/Unoriented/Affine.lean | 2 +- Mathlib/Geometry/Euclidean/PerpBisector.lean | 2 +- Mathlib/RingTheory/AlgebraicIndependent.lean | 6 ++---- 18 files changed, 17 insertions(+), 32 deletions(-) diff --git a/Mathlib/Data/Matrix/Hadamard.lean b/Mathlib/Data/Matrix/Hadamard.lean index 65b6312faeade..da40578d31dd9 100644 --- a/Mathlib/Data/Matrix/Hadamard.lean +++ b/Mathlib/Data/Matrix/Hadamard.lean @@ -30,8 +30,7 @@ hadamard product, hadamard -/ -variable {α β γ m n : Type*} -variable {R : Type*} +variable {α m n R : Type*} namespace Matrix @@ -124,7 +123,7 @@ end Diagonal section trace variable [Fintype m] [Fintype n] -variable (R) [Semiring α] [Semiring R] [Module R α] +variable (R) [Semiring α] theorem sum_hadamard_eq : (∑ i : m, ∑ j : n, (A ⊙ B) i j) = trace (A * Bᵀ) := rfl diff --git a/Mathlib/Data/Matrix/Reflection.lean b/Mathlib/Data/Matrix/Reflection.lean index 222e7f76799ec..54ccf4ce38ed2 100644 --- a/Mathlib/Data/Matrix/Reflection.lean +++ b/Mathlib/Data/Matrix/Reflection.lean @@ -36,7 +36,7 @@ open Matrix namespace Matrix -variable {l m n : ℕ} {α β : Type*} +variable {l m n : ℕ} {α : Type*} /-- `∀` with better defeq for `∀ x : Matrix (Fin m) (Fin n) α, P x`. -/ def Forall : ∀ {m n} (_ : Matrix (Fin m) (Fin n) α → Prop), Prop diff --git a/Mathlib/Data/Matroid/Closure.lean b/Mathlib/Data/Matroid/Closure.lean index 0ae2d18304977..a788bf4ba0f3d 100644 --- a/Mathlib/Data/Matroid/Closure.lean +++ b/Mathlib/Data/Matroid/Closure.lean @@ -274,7 +274,7 @@ lemma mem_closure_self (M : Matroid α) (e : α) (he : e ∈ M.E := by aesop_mat section Indep -variable {ι : Sort*} {I J B : Set α} {f x y : α} +variable {ι : Sort*} {I J B : Set α} {f x : α} lemma Indep.closure_eq_setOf_basis_insert (hI : M.Indep I) : M.closure I = {x | M.Basis I (insert x I)} := by diff --git a/Mathlib/Data/Matroid/IndepAxioms.lean b/Mathlib/Data/Matroid/IndepAxioms.lean index 74e8d39c7eafb..a1769ebc4fc2f 100644 --- a/Mathlib/Data/Matroid/IndepAxioms.lean +++ b/Mathlib/Data/Matroid/IndepAxioms.lean @@ -82,7 +82,7 @@ for the inverse of `e`). open Set Matroid -variable {α : Type*} {I B X : Set α} +variable {α : Type*} section IndepMatroid diff --git a/Mathlib/Data/Matroid/Map.lean b/Mathlib/Data/Matroid/Map.lean index 4cd4e5646e477..2f96ce2ff696f 100644 --- a/Mathlib/Data/Matroid/Map.lean +++ b/Mathlib/Data/Matroid/Map.lean @@ -103,7 +103,7 @@ For this reason, `Matroid.map` requires injectivity to be well-defined in genera open Set Function Set.Notation namespace Matroid -variable {α β : Type*} {f : α → β} {E I s : Set α} {M : Matroid α} {N : Matroid β} +variable {α β : Type*} {f : α → β} {E I : Set α} {M : Matroid α} {N : Matroid β} section comap diff --git a/Mathlib/Data/Matroid/Restrict.lean b/Mathlib/Data/Matroid/Restrict.lean index 3e9f613bea804..8b0954f2e1d4a 100644 --- a/Mathlib/Data/Matroid/Restrict.lean +++ b/Mathlib/Data/Matroid/Restrict.lean @@ -66,7 +66,7 @@ open Set namespace Matroid -variable {α : Type*} {M : Matroid α} {R I J X Y : Set α} +variable {α : Type*} {M : Matroid α} {R I X Y : Set α} section restrict diff --git a/Mathlib/Data/NNRat/BigOperators.lean b/Mathlib/Data/NNRat/BigOperators.lean index 755e9be98bc81..547de4bd5a2e7 100644 --- a/Mathlib/Data/NNRat/BigOperators.lean +++ b/Mathlib/Data/NNRat/BigOperators.lean @@ -10,7 +10,7 @@ import Mathlib.Data.NNRat.Defs /-! # Casting lemmas for non-negative rational numbers involving sums and products -/ -variable {ι α : Type*} +variable {α : Type*} namespace NNRat diff --git a/Mathlib/Data/Nat/Factorial/SuperFactorial.lean b/Mathlib/Data/Nat/Factorial/SuperFactorial.lean index ce246432d96d0..d4ce6e34adfbb 100644 --- a/Mathlib/Data/Nat/Factorial/SuperFactorial.lean +++ b/Mathlib/Data/Nat/Factorial/SuperFactorial.lean @@ -33,8 +33,6 @@ scoped notation "sf" n:60 => Nat.superFactorial n section SuperFactorial -variable {n : ℕ} - @[simp] theorem superFactorial_zero : sf 0 = 1 := rfl diff --git a/Mathlib/Data/Nat/Fib/Zeckendorf.lean b/Mathlib/Data/Nat/Fib/Zeckendorf.lean index 9d1209d75fc3e..ea2b791ad22c0 100644 --- a/Mathlib/Data/Nat/Fib/Zeckendorf.lean +++ b/Mathlib/Data/Nat/Fib/Zeckendorf.lean @@ -67,7 +67,7 @@ lemma IsZeckendorfRep.sum_fib_lt : ∀ {n l}, IsZeckendorfRep l → (∀ a ∈ ( end List namespace Nat -variable {l : List ℕ} {a m n : ℕ} +variable {m n : ℕ} /-- The greatest index of a Fibonacci number less than or equal to `n`. -/ def greatestFib (n : ℕ) : ℕ := (n + 1).findGreatest (fun k ↦ fib k ≤ n) diff --git a/Mathlib/Data/Sum/Interval.lean b/Mathlib/Data/Sum/Interval.lean index 75478d551bcc8..851ec0a28c334 100644 --- a/Mathlib/Data/Sum/Interval.lean +++ b/Mathlib/Data/Sum/Interval.lean @@ -222,7 +222,7 @@ instance instLocallyFiniteOrder : LocallyFiniteOrder (α ⊕ β) where finset_mem_Ioc := by rintro (a | a) (b | b) (x | x) <;> simp finset_mem_Ioo := by rintro (a | a) (b | b) (x | x) <;> simp -variable (a₁ a₂ : α) (b₁ b₂ : β) (a b : α ⊕ β) +variable (a₁ a₂ : α) (b₁ b₂ : β) theorem Icc_inl_inl : Icc (inl a₁ : α ⊕ β) (inl a₂) = (Icc a₁ a₂).map Embedding.inl := rfl diff --git a/Mathlib/Deprecated/Subgroup.lean b/Mathlib/Deprecated/Subgroup.lean index 40cdc4a7f736d..5c168b0eb1a99 100644 --- a/Mathlib/Deprecated/Subgroup.lean +++ b/Mathlib/Deprecated/Subgroup.lean @@ -32,7 +32,7 @@ subgroup, subgroups, IsSubgroup open Set Function -variable {G : Type*} {H : Type*} {A : Type*} {a a₁ a₂ b c : G} +variable {G : Type*} {H : Type*} {A : Type*} {a b : G} section Group diff --git a/Mathlib/Deprecated/Subring.lean b/Mathlib/Deprecated/Subring.lean index 20fc691ea4b93..b4b7c02b787e0 100644 --- a/Mathlib/Deprecated/Subring.lean +++ b/Mathlib/Deprecated/Subring.lean @@ -64,8 +64,6 @@ theorem isSubring_set_range {R : Type u} {S : Type v} [Ring R] [Ring S] (f : R end RingHom -variable {cR : Type u} [CommRing cR] - theorem IsSubring.inter {S₁ S₂ : Set R} (hS₁ : IsSubring S₁) (hS₂ : IsSubring S₂) : IsSubring (S₁ ∩ S₂) := { IsAddSubgroup.inter hS₁.toIsAddSubgroup hS₂.toIsAddSubgroup, diff --git a/Mathlib/FieldTheory/Finite/Polynomial.lean b/Mathlib/FieldTheory/Finite/Polynomial.lean index 36dd816621f74..fa34749883456 100644 --- a/Mathlib/FieldTheory/Finite/Polynomial.lean +++ b/Mathlib/FieldTheory/Finite/Polynomial.lean @@ -173,18 +173,12 @@ noncomputable instance [CommRing K] : Inhabited (R σ K) := def evalᵢ [CommRing K] : R σ K →ₗ[K] (σ → K) → K := (evalₗ K σ).comp (restrictDegree σ K (Fintype.card K - 1)).subtype -section CommRing - -variable [CommRing K] - -- TODO: would be nice to replace this by suitable decidability assumptions open Classical in noncomputable instance decidableRestrictDegree (m : ℕ) : DecidablePred (· ∈ { n : σ →₀ ℕ | ∀ i, n i ≤ m }) := by simp only [Set.mem_setOf_eq]; infer_instance -end CommRing - variable [Field K] open Classical in diff --git a/Mathlib/FieldTheory/IsAlgClosed/Classification.lean b/Mathlib/FieldTheory/IsAlgClosed/Classification.lean index 5d028a07e9d41..99b41accb1cf2 100644 --- a/Mathlib/FieldTheory/IsAlgClosed/Classification.lean +++ b/Mathlib/FieldTheory/IsAlgClosed/Classification.lean @@ -119,10 +119,8 @@ end Classification section Cardinal -variable {R L K : Type u} [CommRing R] -variable [Field K] [Algebra R K] [IsAlgClosed K] +variable {R K : Type u} [CommRing R] [Field K] [Algebra R K] [IsAlgClosed K] variable {ι : Type u} (v : ι → K) -variable (hv : IsTranscendenceBasis R v) theorem cardinal_le_max_transcendence_basis (hv : IsTranscendenceBasis R v) : #K ≤ max (max #R #ι) ℵ₀ := diff --git a/Mathlib/FieldTheory/IsPerfectClosure.lean b/Mathlib/FieldTheory/IsPerfectClosure.lean index 743199124c340..5c6d7928dc1a4 100644 --- a/Mathlib/FieldTheory/IsPerfectClosure.lean +++ b/Mathlib/FieldTheory/IsPerfectClosure.lean @@ -196,7 +196,7 @@ theorem RingHom.pNilradical_le_ker_of_perfectRing [ExpChar L p] [PerfectRing L p rwa [map_pow, ← iterateFrobenius_def, ← iterateFrobeniusEquiv_apply, RingEquiv.symm_apply_apply, map_zero, map_zero] at h -variable [ExpChar L p] [ExpChar M p] in +variable [ExpChar L p] in theorem IsPerfectClosure.ker_eq [PerfectRing L p] [IsPerfectClosure i p] : RingHom.ker i = pNilradical K p := IsPRadical.ker_le'.antisymm (i.pNilradical_le_ker_of_perfectRing p) diff --git a/Mathlib/Geometry/Euclidean/Angle/Unoriented/Affine.lean b/Mathlib/Geometry/Euclidean/Angle/Unoriented/Affine.lean index b58876fe73b9d..36b42892368e8 100644 --- a/Mathlib/Geometry/Euclidean/Angle/Unoriented/Affine.lean +++ b/Mathlib/Geometry/Euclidean/Angle/Unoriented/Affine.lean @@ -33,7 +33,7 @@ namespace EuclideanGeometry open InnerProductGeometry variable {V P : Type*} [NormedAddCommGroup V] [InnerProductSpace ℝ V] [MetricSpace P] - [NormedAddTorsor V P] {p p₀ p₁ p₂ : P} + [NormedAddTorsor V P] {p p₀ : P} /-- The undirected angle at `p2` between the line segments to `p1` and `p3`. If either of those points equals `p2`, this is π/2. Use diff --git a/Mathlib/Geometry/Euclidean/PerpBisector.lean b/Mathlib/Geometry/Euclidean/PerpBisector.lean index 00730a11e550c..b2d97cbe87082 100644 --- a/Mathlib/Geometry/Euclidean/PerpBisector.lean +++ b/Mathlib/Geometry/Euclidean/PerpBisector.lean @@ -29,7 +29,7 @@ noncomputable section namespace AffineSubspace -variable {c c₁ c₂ p₁ p₂ : P} +variable {c p₁ p₂ : P} /-- Perpendicular bisector of a segment in a Euclidean affine space. -/ def perpBisector (p₁ p₂ : P) : AffineSubspace ℝ P := diff --git a/Mathlib/RingTheory/AlgebraicIndependent.lean b/Mathlib/RingTheory/AlgebraicIndependent.lean index 0e358adaa9418..43e4a7ea67e93 100644 --- a/Mathlib/RingTheory/AlgebraicIndependent.lean +++ b/Mathlib/RingTheory/AlgebraicIndependent.lean @@ -47,11 +47,9 @@ open scoped Classical universe x u v w variable {ι : Type*} {ι' : Type*} (R : Type*) {K : Type*} -variable {A : Type*} {A' A'' : Type*} {V : Type u} {V' : Type*} +variable {A : Type*} {A' : Type*} variable (x : ι → A) -variable [CommRing R] [CommRing A] [CommRing A'] [CommRing A''] -variable [Algebra R A] [Algebra R A'] [Algebra R A''] -variable {a b : R} +variable [CommRing R] [CommRing A] [CommRing A'] [Algebra R A] [Algebra R A'] /-- `AlgebraicIndependent R x` states the family of elements `x` is algebraically independent over `R`, meaning that the canonical From ff1b88672d4e66db5195c5c57684e4081fab209f Mon Sep 17 00:00:00 2001 From: Matthew Robert Ballard Date: Mon, 21 Oct 2024 09:56:18 +0000 Subject: [PATCH 233/505] chore: add test to keep powering in `ZMod` fast (#17996) Thanks to #8885 we can now evaluate powers in `ZMod M` much more efficiently for large `M`. To guard against reversion, we add a test that for a chosen 1024-bit prime and base we should satisfy Fermat's Little Theorem when `#eval` is used. --- test/BinPow.lean | 13 +++++++++++++ 1 file changed, 13 insertions(+) create mode 100644 test/BinPow.lean diff --git a/test/BinPow.lean b/test/BinPow.lean new file mode 100644 index 0000000000000..33b61a251e7fd --- /dev/null +++ b/test/BinPow.lean @@ -0,0 +1,13 @@ +import Mathlib.Data.ZMod.Defs + +-- A 1024-bit prime number +set_option linter.style.longLine false in +abbrev M : Nat := 0xb10b8f96a080e01dde92de5eae5d54ec52c99fbcfb06a3c69a6a9dca52d23b616073e28675a23d189838ef1e2ee652c013ecb4aea906112324975c3cd49b83bfaccbdd7d90c4bd7098488e9c219a73724effd6fae5644738faa31a4ff55bccc0a151af5f0dc8b4bd45bf37df365c1a65e68cfda76d4da708df1fb2bc2e4a4371 + +set_option linter.style.longLine false in +abbrev g : Nat := 0xa4d1cbd5c3fd34126765a442efb99905f8104dd258ac507fd6406cff14266d31266fea1e5c41564b777e690f5504f213160217b4b01b886a5e91547f9e2749f4d7fbd7d3b9a92ee1909d0d2263f80a76a6a24c087a091f531dbf0a0169b6a28ad662a4d18e73afa32d779d5918d08bc8858f4dcef97c2a24855e6eeb22b3b2e5 + +-- This should evaluate within a few seconds compared to never(ish) previously +/-- info: 1 -/ +#guard_msgs in +#eval (g : ZMod M) ^ (M - 1) From 1df44a89ac2cd5c83c4b6ea19b2aa5ed3f99c479 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Mon, 21 Oct 2024 10:04:43 +0000 Subject: [PATCH 234/505] chore: remove `CoeFun` instances where `FunLike` is available (#17917) During the port we found that `FunLike` is robust enough not to need an extra `CoeFun` shortcut. Let's see if that rule can be consistently applied to the whole of the library. There is still duplication between `FunLike` and `CoeFun` for `Grp`, `Mon`, `CommGrp` and `CommMon`, which will need a more thorough fix. See also #17866. This contains everything except the interval [1/4, 3/8] of the changes, and doesn't affect the benchmark. --- Mathlib/LinearAlgebra/AffineSpace/ContinuousAffineEquiv.lean | 3 --- Mathlib/LinearAlgebra/Alternating/Basic.lean | 4 ---- 2 files changed, 7 deletions(-) diff --git a/Mathlib/LinearAlgebra/AffineSpace/ContinuousAffineEquiv.lean b/Mathlib/LinearAlgebra/AffineSpace/ContinuousAffineEquiv.lean index 12ffbd6ace838..df744660c3393 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/ContinuousAffineEquiv.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/ContinuousAffineEquiv.lean @@ -73,9 +73,6 @@ instance instEquivLike : EquivLike (P₁ ≃ᵃL[k] P₂) P₁ P₂ where right_inv f := f.right_inv coe_injective' _ _ h _ := toAffineEquiv_injective (DFunLike.coe_injective h) -instance : CoeFun (P₁ ≃ᵃL[k] P₂) fun _ ↦ P₁ → P₂ := - DFunLike.hasCoeToFun - attribute [coe] ContinuousAffineEquiv.toAffineEquiv /-- Coerce continuous affine equivalences to affine equivalences. -/ diff --git a/Mathlib/LinearAlgebra/Alternating/Basic.lean b/Mathlib/LinearAlgebra/Alternating/Basic.lean index f3a7aafc8973a..06184ef9cb523 100644 --- a/Mathlib/LinearAlgebra/Alternating/Basic.lean +++ b/Mathlib/LinearAlgebra/Alternating/Basic.lean @@ -93,10 +93,6 @@ instance instFunLike : FunLike (M [⋀^ι]→ₗ[R] N) (ι → M) N where rcases g with ⟨⟨_, _, _⟩, _⟩ congr --- shortcut instance -instance coeFun : CoeFun (M [⋀^ι]→ₗ[R] N) fun _ => (ι → M) → N := - ⟨DFunLike.coe⟩ - initialize_simps_projections AlternatingMap (toFun → apply) @[simp] From d43bc00b660d5456721455c8f4e4d1d2244f555b Mon Sep 17 00:00:00 2001 From: tukamilano Date: Mon, 21 Oct 2024 10:47:33 +0000 Subject: [PATCH 235/505] feat: Popoviciu's inequality (#17868) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Formalize [Popoviciu's inequality on variances](https://en.wikipedia.org/wiki/Popoviciu%27s_inequality_on_variances) through the [Bhatia-Davis inequality](https://en.wikipedia.org/wiki/Bhatia%E2%80%93Davis_inequality). Co-authored-by: Yuma Mizuno mizuno.y.aj@gmail.com Co-authored-by: tukamilano <107251402+tukamilano@users.noreply.github.com> Co-authored-by: Yaël Dillies --- .../Function/LpSeminorm/Basic.lean | 7 +++ Mathlib/Probability/Variance.lean | 50 ++++++++++++++++++- 2 files changed, 56 insertions(+), 1 deletion(-) diff --git a/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean b/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean index ee75398764b6e..b7dc71cd9a9e8 100644 --- a/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean +++ b/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean @@ -721,6 +721,13 @@ theorem Memℒp.of_bound [IsFiniteMeasure μ] {f : α → E} (hf : AEStronglyMea (hfC : ∀ᵐ x ∂μ, ‖f x‖ ≤ C) : Memℒp f p μ := (memℒp_const C).of_le hf (hfC.mono fun _x hx => le_trans hx (le_abs_self _)) +theorem memℒp_of_bounded [IsFiniteMeasure μ] + {a b : ℝ} {f : α → ℝ} (h : ∀ᵐ x ∂μ, f x ∈ Set.Icc a b) + (hX : AEStronglyMeasurable f μ) (p : ENNReal) : Memℒp f p μ := + have ha : ∀ᵐ x ∂μ, a ≤ f x := h.mono fun ω h => h.1 + have hb : ∀ᵐ x ∂μ, f x ≤ b := h.mono fun ω h => h.2 + (memℒp_const (max |a| |b|)).mono' hX (by filter_upwards [ha, hb] with x using abs_le_max_abs_abs) + @[mono] theorem eLpNorm'_mono_measure (f : α → F) (hμν : ν ≤ μ) (hq : 0 ≤ q) : eLpNorm' f q ν ≤ eLpNorm' f q μ := by diff --git a/Mathlib/Probability/Variance.lean b/Mathlib/Probability/Variance.lean index 3510fdd559901..00263fc96ac34 100644 --- a/Mathlib/Probability/Variance.lean +++ b/Mathlib/Probability/Variance.lean @@ -30,9 +30,12 @@ We define the variance of a real-valued random variable as `Var[X] = 𝔼[(X - random variables is the sum of the variances. * `ProbabilityTheory.IndepFun.variance_sum`: the variance of a finite sum of pairwise independent random variables is the sum of the variances. +* `ProbabilityTheory.variance_le_sub_mul_sub`: the variance of a random variable `X` satisfying + `a ≤ X ≤ b` almost everywhere is at most `(b - 𝔼 X) * (𝔼 X - a)`. +* `ProbabilityTheory.variance_le_sq_of_bounded`: the variance of a random variable `X` satisfying + `a ≤ X ≤ b` almost everywhere is at most`((b - a) / 2) ^ 2`. -/ - open MeasureTheory Filter Finset noncomputable section @@ -341,4 +344,49 @@ theorem IndepFun.variance_sum [IsProbabilityMeasure μ] {ι : Type*} {X : ι → rw [IH (fun i hi => hs i (mem_insert_of_mem hi)) (h.mono (by simp only [coe_insert, Set.subset_insert]))] +/-- **The Bhatia-Davis inequality on variance** + +The variance of a random variable `X` satisfying `a ≤ X ≤ b` almost everywhere is at most +`(b - 𝔼 X) * (𝔼 X - a)`. -/ +lemma variance_le_sub_mul_sub [IsProbabilityMeasure μ] {a b : ℝ} {X : Ω → ℝ} + (h : ∀ᵐ ω ∂μ, X ω ∈ Set.Icc a b) (hX : AEMeasurable X μ) : + variance X μ ≤ (b - μ[X]) * (μ[X] - a) := by + have ha : ∀ᵐ ω ∂μ, a ≤ X ω := h.mono fun ω h => h.1 + have hb : ∀ᵐ ω ∂μ, X ω ≤ b := h.mono fun ω h => h.2 + have hX_int₂ : Integrable (fun ω ↦ -X ω ^ 2) μ := + (memℒp_of_bounded h hX.aestronglyMeasurable 2).integrable_sq.neg + have hX_int₁ : Integrable (fun ω ↦ (a + b) * X ω) μ := + ((integrable_const (max |a| |b|)).mono' hX.aestronglyMeasurable + (by filter_upwards [ha, hb] with ω using abs_le_max_abs_abs)).const_mul (a + b) + have h0 : 0 ≤ - μ[X ^ 2] + (a + b) * μ[X] - a * b := + calc + _ ≤ ∫ ω, (b - X ω) * (X ω - a) ∂μ := by + apply integral_nonneg_of_ae + filter_upwards [ha, hb] with ω ha' hb' + exact mul_nonneg (by linarith : 0 ≤ b - X ω) (by linarith : 0 ≤ X ω - a) + _ = ∫ ω, - X ω ^ 2 + (a + b) * X ω - a * b ∂μ := + integral_congr_ae <| ae_of_all μ fun ω ↦ by ring + _ = ∫ ω, - X ω ^ 2 + (a + b) * X ω ∂μ - ∫ _, a * b ∂μ := + integral_sub (hX_int₂.add hX_int₁) (integrable_const (a * b)) + _ = ∫ ω, - X ω ^ 2 + (a + b) * X ω ∂μ - a * b := by simp + _ = - μ[X ^ 2] + (a + b) * μ[X] - a * b := by + simp [← integral_neg, ← integral_mul_left, integral_add hX_int₂ hX_int₁] + calc + _ ≤ (a + b) * μ[X] - a * b - μ[X] ^ 2 := by + rw [variance_def' (memℒp_of_bounded h hX.aestronglyMeasurable 2)] + linarith + _ = (b - μ[X]) * (μ[X] - a) := by ring + +/-- **Popoviciu's inequality on variance** + +The variance of a random variable `X` satisfying `a ≤ X ≤ b` almost everywhere is at most +`((b - a) / 2) ^ 2`. -/ +lemma variance_le_sq_of_bounded [IsProbabilityMeasure μ] {a b : ℝ} {X : Ω → ℝ} + (h : ∀ᵐ ω ∂μ, X ω ∈ Set.Icc a b) (hX : AEMeasurable X μ) : + variance X μ ≤ ((b - a) / 2) ^ 2 := + calc + _ ≤ (b - μ[X]) * (μ[X] - a) := variance_le_sub_mul_sub h hX + _ = ((b - a) / 2) ^ 2 - (μ[X] - (b + a) / 2) ^ 2 := by ring + _ ≤ ((b - a) / 2) ^ 2 := sub_le_self _ (sq_nonneg _) + end ProbabilityTheory From 3bf558f968bd0f2519e2c605638a91a2f84f65c5 Mon Sep 17 00:00:00 2001 From: Xavier Roblot Date: Mon, 21 Oct 2024 13:21:29 +0000 Subject: [PATCH 236/505] feat(Module/ZLattice): define the pullback of a ZLattice (#16822) Define the pullback of a ZLattice `L` by a linear map `f`. The following results are also included: - If `f` is a continuous linear equiv, then the pullback is also a ZLattice (added as an instance). - If `f` is a linear equiv, define the corresponding basis of the pullback obtained from a basis of `L` - If `f` is a continuous linear equiv and volume preserving, prove that `L` and its pullback have the same volume. This PR is part of the proof of the Analytic Class Number Formula. Co-authored-by: Xavier Roblot <46200072+xroblot@users.noreply.github.com> --- Mathlib/Algebra/Module/ZLattice/Basic.lean | 89 ++++++++++++++++++- Mathlib/Algebra/Module/ZLattice/Covolume.lean | 28 ++++-- 2 files changed, 108 insertions(+), 9 deletions(-) diff --git a/Mathlib/Algebra/Module/ZLattice/Basic.lean b/Mathlib/Algebra/Module/ZLattice/Basic.lean index f5d850e0a404e..8113915f8a6a5 100644 --- a/Mathlib/Algebra/Module/ZLattice/Basic.lean +++ b/Mathlib/Algebra/Module/ZLattice/Basic.lean @@ -24,7 +24,7 @@ A `ℤ`-lattice `L` can be defined in two ways: Results about the first point of view are in the `ZSpan` namespace and results about the second point of view are in the `ZLattice` namespace. -## Main results +## Main results and definitions * `ZSpan.isAddFundamentalDomain`: for a ℤ-lattice `Submodule.span ℤ (Set.range b)`, proves that the set defined by `ZSpan.fundamentalDomain` is a fundamental domain. @@ -32,6 +32,9 @@ the set defined by `ZSpan.fundamentalDomain` is a fundamental domain. `ℤ`-module * `ZLattice.rank`: a `ℤ`-submodule of `E` that is discrete and spans `E` over `K` is free of `ℤ`-rank equal to the `K`-rank of `E` +* `ZLattice.comap`: for `e : E → F` a linear map and `L : Submodule ℤ E`, define the pullback of +`L` by `e`. If `L` is a `IsZLattice` and `e` is a continuous linear equiv, then it is also a +`IsZLattice`, see `instIsZLatticeComap`. ## Implementation Notes @@ -58,6 +61,11 @@ variable (b : Basis ι K E) theorem span_top : span K (span ℤ (Set.range b) : Set E) = ⊤ := by simp [span_span_of_tower] + +theorem map {F : Type*} [NormedAddCommGroup F] [NormedSpace K F] (f : E ≃ₗ[K] F) : + Submodule.map (f.restrictScalars ℤ) (span ℤ (Set.range b)) = span ℤ (Set.range (b.map f)) := by + simp_rw [Submodule.map_span, LinearEquiv.restrictScalars_apply, Basis.coe_map, Set.range_comp] + /-- The fundamental domain of the ℤ-lattice spanned by `b`. See `ZSpan.isAddFundamentalDomain` for the proof that it is a fundamental domain. -/ def fundamentalDomain : Set E := {m | ∀ i, b.repr m i ∈ Set.Ico (0 : K) 1} @@ -399,6 +407,8 @@ theorem _root_.ZSpan.isZLattice {E ι : Type*} [NormedAddCommGroup E] [NormedSpa IsZLattice ℝ (span ℤ (Set.range b)) where span_top := ZSpan.span_top b +section NormedLinearOrderedField + variable (K : Type*) [NormedLinearOrderedField K] [HasSolidNorm K] [FloorRing K] variable {E : Type*} [NormedAddCommGroup E] [NormedSpace K E] [FiniteDimensional K E] variable [ProperSpace E] (L : Submodule ℤ E) [DiscreteTopology L] @@ -607,4 +617,81 @@ instance instCountable_of_discrete_submodule {E : Type*} [NormedAddCommGroup E] simp_rw [← (Module.Free.chooseBasis ℤ L).ofZLatticeBasis_span ℝ] infer_instance +end NormedLinearOrderedField + +section comap + +variable (K : Type*) [NormedField K] {E F : Type*} [NormedAddCommGroup E] [NormedSpace K E] + [NormedAddCommGroup F] [NormedSpace K F] (L : Submodule ℤ E) + +/-- Let `e : E → F` a linear map, the map that sends a `L : Submodule ℤ E` to the +`Submodule ℤ F` that is the pullback of `L` by `e`. If `IsZLattice L` and `e` is a continuous +linear equiv, then it is a `IsZLattice` of `E`, see `instIsZLatticeComap`. -/ +protected def ZLattice.comap (e : F →ₗ[K] E) := L.comap (e.restrictScalars ℤ) + +@[simp] +theorem ZLattice.coe_comap (e : F →ₗ[K] E) : + (ZLattice.comap K L e : Set F) = e⁻¹' L := rfl + +theorem ZLattice.comap_refl : + ZLattice.comap K L (1 : E →ₗ[K] E)= L := Submodule.comap_id L + +theorem ZLattice.comap_discreteTopology [hL : DiscreteTopology L] {e : F →ₗ[K] E} + (he₁ : Continuous e) (he₂ : Function.Injective e) : + DiscreteTopology (ZLattice.comap K L e) := by + exact DiscreteTopology.preimage_of_continuous_injective L he₁ he₂ + +instance [DiscreteTopology L] (e : F ≃L[K] E) : + DiscreteTopology (ZLattice.comap K L e.toLinearMap) := + ZLattice.comap_discreteTopology K L e.continuous e.injective + +theorem ZLattice.comap_span_top (hL : span K (L : Set E) = ⊤) {e : F →ₗ[K] E} + (he : (L : Set E) ⊆ LinearMap.range e) : + span K (ZLattice.comap K L e : Set F) = ⊤ := by + rw [ZLattice.coe_comap, Submodule.span_preimage_eq (Submodule.nonempty L) he, hL, comap_top] + +instance instIsZLatticeComap [DiscreteTopology L] [IsZLattice K L] (e : F ≃L[K] E) : + IsZLattice K (ZLattice.comap K L e.toLinearMap) where + span_top := by + rw [ZLattice.coe_comap, LinearEquiv.coe_coe, e.coe_toLinearEquiv, ← e.image_symm_eq_preimage, + ← Submodule.map_span, IsZLattice.span_top, Submodule.map_top, LinearEquivClass.range] + +theorem ZLattice.comap_comp {G : Type*} [NormedAddCommGroup G] [NormedSpace K G] + (e : F →ₗ[K] E) (e' : G →ₗ[K] F) : + (ZLattice.comap K (ZLattice.comap K L e) e') = ZLattice.comap K L (e ∘ₗ e') := + (Submodule.comap_comp _ _ L).symm + +/-- If `e` is a linear equivalence, it induces a `ℤ`-linear equivalence between +`L` and `ZLattice.comap K L e`. -/ +def ZLattice.comap_equiv (e : F ≃ₗ[K] E) : + L ≃ₗ[ℤ] (ZLattice.comap K L e.toLinearMap) := + LinearEquiv.ofBijective + ((e.symm.toLinearMap.restrictScalars ℤ).restrict + (fun _ h ↦ by simpa [← SetLike.mem_coe] using h)) + ⟨fun _ _ h ↦ Subtype.ext_iff_val.mpr (e.symm.injective (congr_arg Subtype.val h)), + fun ⟨x, hx⟩ ↦ ⟨⟨e x, by rwa [← SetLike.mem_coe, ZLattice.coe_comap] at hx⟩, + by simp [Subtype.ext_iff_val]⟩⟩ + +@[simp] +theorem ZLattice.comap_equiv_apply (e : F ≃ₗ[K] E) (x : L) : + ZLattice.comap_equiv K L e x = e.symm x := rfl + +/-- The basis of `ZLattice.comap K L e` given by the image of a basis `b` of `L` by `e.symm`. -/ +def Basis.ofZLatticeComap (e : F ≃ₗ[K] E) {ι : Type*} + (b : Basis ι ℤ L) : + Basis ι ℤ (ZLattice.comap K L e.toLinearMap) := b.map (ZLattice.comap_equiv K L e) + +@[simp] +theorem Basis.ofZLatticeComap_apply (e : F ≃ₗ[K] E) {ι : Type*} + (b : Basis ι ℤ L) (i : ι) : + b.ofZLatticeComap K L e i = e.symm (b i) := by simp [Basis.ofZLatticeComap] + +@[simp] +theorem Basis.ofZLatticeComap_repr_apply (e : F ≃ₗ[K] E) {ι : Type*} (b : Basis ι ℤ L) (x : L) + (i : ι) : + (b.ofZLatticeComap K L e).repr (ZLattice.comap_equiv K L e x) i = b.repr x i := by + simp [Basis.ofZLatticeComap] + +end comap + end ZLattice diff --git a/Mathlib/Algebra/Module/ZLattice/Covolume.lean b/Mathlib/Algebra/Module/ZLattice/Covolume.lean index 04e559d80a384..8b3705ee2aa1e 100644 --- a/Mathlib/Algebra/Module/ZLattice/Covolume.lean +++ b/Mathlib/Algebra/Module/ZLattice/Covolume.lean @@ -8,7 +8,7 @@ import Mathlib.Algebra.Module.ZLattice.Basic /-! # Covolume of ℤ-lattices -Let `E` be a finite dimensional real vector space with an inner product. +Let `E` be a finite dimensional real vector space. Let `L` be a `ℤ`-lattice `L` defined as a discrete `ℤ`-submodule of `E` that spans `E` over `ℝ`. @@ -29,7 +29,7 @@ noncomputable section namespace ZLattice -open Submodule MeasureTheory Module MeasureTheory Module +open Submodule MeasureTheory Module MeasureTheory Module ZSpan section General @@ -61,18 +61,30 @@ theorem covolume_eq_measure_fundamentalDomain {F : Set E} (h : IsAddFundamentalD theorem covolume_ne_zero : covolume L μ ≠ 0 := by rw [covolume_eq_measure_fundamentalDomain L μ (isAddFundamentalDomain (Free.chooseBasis ℤ L) μ), ENNReal.toReal_ne_zero] - refine ⟨ZSpan.measure_fundamentalDomain_ne_zero _, ne_of_lt ?_⟩ - exact Bornology.IsBounded.measure_lt_top (ZSpan.fundamentalDomain_isBounded _) + refine ⟨measure_fundamentalDomain_ne_zero _, ne_of_lt ?_⟩ + exact Bornology.IsBounded.measure_lt_top (fundamentalDomain_isBounded _) theorem covolume_pos : 0 < covolume L μ := lt_of_le_of_ne ENNReal.toReal_nonneg (covolume_ne_zero L μ).symm +theorem covolume_comap {F : Type*} [NormedAddCommGroup F] [NormedSpace ℝ F] [FiniteDimensional ℝ F] + [MeasurableSpace F] [BorelSpace F] (ν : Measure F := by volume_tac) [Measure.IsAddHaarMeasure ν] + {e : F ≃L[ℝ] E} (he : MeasurePreserving e ν μ) : + covolume (ZLattice.comap ℝ L e.toLinearMap) ν = covolume L μ := by + rw [covolume_eq_measure_fundamentalDomain _ _ (isAddFundamentalDomain (Free.chooseBasis ℤ L) μ), + covolume_eq_measure_fundamentalDomain _ _ ((isAddFundamentalDomain + ((Free.chooseBasis ℤ L).ofZLatticeComap ℝ L e.toLinearEquiv) ν)), ← he.measure_preimage + (fundamentalDomain_measurableSet _).nullMeasurableSet, ← e.image_symm_eq_preimage, + ← e.symm.coe_toLinearEquiv, map_fundamentalDomain] + congr! + ext; simp + theorem covolume_eq_det_mul_measure {ι : Type*} [Fintype ι] [DecidableEq ι] (b : Basis ι ℤ L) (b₀ : Basis ι ℝ E) : - covolume L μ = |b₀.det ((↑) ∘ b)| * (μ (ZSpan.fundamentalDomain b₀)).toReal := by + covolume L μ = |b₀.det ((↑) ∘ b)| * (μ (fundamentalDomain b₀)).toReal := by rw [covolume_eq_measure_fundamentalDomain L μ (isAddFundamentalDomain b μ), - ZSpan.measure_fundamentalDomain _ _ b₀, - measure_congr (ZSpan.fundamentalDomain_ae_parallelepiped b₀ μ), ENNReal.toReal_mul, + measure_fundamentalDomain _ _ b₀, + measure_congr (fundamentalDomain_ae_parallelepiped b₀ μ), ENNReal.toReal_mul, ENNReal.toReal_ofReal (by positivity)] congr ext @@ -82,7 +94,7 @@ theorem covolume_eq_det {ι : Type*} [Fintype ι] [DecidableEq ι] (L : Submodul [DiscreteTopology L] [IsZLattice ℝ L] (b : Basis ι ℤ L) : covolume L = |(Matrix.of ((↑) ∘ b)).det| := by rw [covolume_eq_measure_fundamentalDomain L volume (isAddFundamentalDomain b volume), - ZSpan.volume_fundamentalDomain, ENNReal.toReal_ofReal (by positivity)] + volume_fundamentalDomain, ENNReal.toReal_ofReal (by positivity)] congr ext1 exact b.ofZLatticeBasis_apply ℝ L _ From f4c1fd5536832599eba6841ba34af18120c8b5ff Mon Sep 17 00:00:00 2001 From: judithludwig Date: Mon, 21 Oct 2024 14:11:35 +0000 Subject: [PATCH 237/505] =?UTF-8?q?feat:=20`a=20^=20n=20=3D=201=20?= =?UTF-8?q?=E2=86=94=20n=20=3D=200`=20and=20similar=20(#17745)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR adds a few lemmas about GroupsWithZero, more precisely: - the lemmas `zero_pow_eq_one₀`, `zero_zpow_eq_one₀` to Algebra.GroupWithZero.Basic and - `zpow_le_zpow_iff_right_of_lt_one₀`, `zpow_lt_zpow_iff_right_of_lt_one₀` and `zpow_eq_one_iff_right₀` to Algebra.Order.GroupWithZero.Unbundled. Used in project Formalization of the Bruhat-Tits Tree. Co-authored-by: Christian Merten Co-authored-by: Christian Merten Co-authored-by: judithludwig <150042763+judithludwig@users.noreply.github.com> --- Mathlib/Algebra/GroupWithZero/Basic.lean | 6 ++++++ .../Algebra/Order/GroupWithZero/Unbundled.lean | 15 +++++++++++++-- 2 files changed, 19 insertions(+), 2 deletions(-) diff --git a/Mathlib/Algebra/GroupWithZero/Basic.lean b/Mathlib/Algebra/GroupWithZero/Basic.lean index 64fdcf4a2b337..c09c2c9f7742f 100644 --- a/Mathlib/Algebra/GroupWithZero/Basic.lean +++ b/Mathlib/Algebra/GroupWithZero/Basic.lean @@ -146,6 +146,9 @@ lemma zero_pow_eq (n : ℕ) : (0 : M₀) ^ n = if n = 0 then 1 else 0 := by · rw [h, pow_zero] · rw [zero_pow h] +lemma zero_pow_eq_one₀ [Nontrivial M₀] : (0 : M₀) ^ n = 1 ↔ n = 0 := by + rw [zero_pow_eq, one_ne_zero.ite_eq_left_iff] + lemma pow_eq_zero_of_le : ∀ {m n} (_ : m ≤ n) (_ : a ^ m = 0), a ^ n = 0 | _, _, Nat.le.refl, ha => ha | _, _, Nat.le.step hmn, ha => by rw [pow_succ, pow_eq_zero_of_le hmn ha, zero_mul] @@ -389,6 +392,9 @@ lemma zero_zpow_eq (n : ℤ) : (0 : G₀) ^ n = if n = 0 then 1 else 0 := by · rw [h, zpow_zero] · rw [zero_zpow _ h] +lemma zero_zpow_eq_one₀ {n : ℤ} : (0 : G₀) ^ n = 1 ↔ n = 0 := by + rw [zero_zpow_eq, one_ne_zero.ite_eq_left_iff] + lemma zpow_add_one₀ (ha : a ≠ 0) : ∀ n : ℤ, a ^ (n + 1) = a ^ n * a | (n : ℕ) => by simp only [← Int.ofNat_succ, zpow_natCast, pow_succ] | .negSucc 0 => by erw [zpow_zero, zpow_negSucc, pow_one, inv_mul_cancel₀ ha] diff --git a/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean b/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean index fbec3165e23ef..2c7df6d747cef 100644 --- a/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean +++ b/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean @@ -1407,8 +1407,8 @@ lemma zpow_lt_zpow_right₀ (ha : 1 < a) (hmn : m < n) : a ^ m < a ^ n := zpow_right_strictMono₀ ha hmn @[gcongr] -lemma zpow_lt_zpow_right_of_lt_one₀ (ha₀ : 0 < a) (ha₁ : a ≤ 1) (hmn : m ≤ n) : a ^ n ≤ a ^ m := - zpow_right_anti₀ ha₀ ha₁ hmn +lemma zpow_lt_zpow_right_of_lt_one₀ (ha₀ : 0 < a) (ha₁ : a < 1) (hmn : m < n) : a ^ n < a ^ m := + zpow_right_strictAnti₀ ha₀ ha₁ hmn lemma one_lt_zpow₀ (ha : 1 < a) (hn : 0 < n) : 1 < a ^ n := by simpa using zpow_right_strictMono₀ ha hn @@ -1428,6 +1428,12 @@ lemma one_lt_zpow_of_neg₀ (ha₀ : 0 < a) (ha₁ : a < 1) (hn : n < 0) : 1 < a @[simp] lemma zpow_lt_zpow_iff_right₀ (ha : 1 < a) : a ^ m < a ^ n ↔ m < n := (zpow_right_strictMono₀ ha).lt_iff_lt +lemma zpow_le_zpow_iff_right_of_lt_one₀ (ha₀ : 0 < a) (ha₁ : a < 1) : + a ^ m ≤ a ^ n ↔ n ≤ m := (zpow_right_strictAnti₀ ha₀ ha₁).le_iff_le + +lemma zpow_lt_zpow_iff_right_of_lt_one₀ (ha₀ : 0 < a) (ha₁ : a < 1) : + a ^ m < a ^ n ↔ n < m := (zpow_right_strictAnti₀ ha₀ ha₁).lt_iff_lt + end PosMulStrictMono section MulPosStrictMono @@ -1512,6 +1518,11 @@ lemma zpow_right_injective₀ (ha₀ : 0 < a) (ha₁ : a ≠ 1) : Injective fun @[simp] lemma zpow_right_inj₀ (ha₀ : 0 < a) (ha₁ : a ≠ 1) : a ^ m = a ^ n ↔ m = n := (zpow_right_injective₀ ha₀ ha₁).eq_iff +lemma zpow_eq_one_iff_right₀ (ha₀ : 0 ≤ a) (ha₁ : a ≠ 1) {n : ℤ} : a ^ n = 1 ↔ n = 0 := by + obtain rfl | ha₀ := ha₀.eq_or_lt + · exact zero_zpow_eq_one₀ + simpa using zpow_right_inj₀ ha₀ ha₁ (n := 0) + end GroupWithZero.LinearOrder section CommSemigroupHasZero From d254991000b87f4ff40cd3647d42c5e3b26c030f Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Mar=C3=ADa=20In=C3=A9s=20de=20Frutos-Fern=C3=A1ndez?= Date: Mon, 21 Oct 2024 14:44:11 +0000 Subject: [PATCH 238/505] feat(Data/Real/IsNonarchimedean): add lemmas (#16767) Used in #15373. Co-authored-by: Yakov Pechersky Co-authored-by: Yakov Pechersky Co-authored-by: David Loeffler --- Mathlib.lean | 3 + Mathlib/Algebra/Order/Hom/Normed.lean | 75 +++++++++++++++ Mathlib/Algebra/Order/Hom/Ultra.lean | 32 +++++++ Mathlib/Analysis/Normed/Group/Ultra.lean | 48 ++++++++++ Mathlib/Data/Real/IsNonarchimedean.lean | 114 +++++++++++++++++++++++ 5 files changed, 272 insertions(+) create mode 100644 Mathlib/Algebra/Order/Hom/Normed.lean create mode 100644 Mathlib/Algebra/Order/Hom/Ultra.lean create mode 100644 Mathlib/Data/Real/IsNonarchimedean.lean diff --git a/Mathlib.lean b/Mathlib.lean index c6e807c4602bf..1f124da7a5e7a 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -627,7 +627,9 @@ import Mathlib.Algebra.Order.GroupWithZero.Unbundled.Lemmas import Mathlib.Algebra.Order.GroupWithZero.WithZero import Mathlib.Algebra.Order.Hom.Basic import Mathlib.Algebra.Order.Hom.Monoid +import Mathlib.Algebra.Order.Hom.Normed import Mathlib.Algebra.Order.Hom.Ring +import Mathlib.Algebra.Order.Hom.Ultra import Mathlib.Algebra.Order.Interval.Basic import Mathlib.Algebra.Order.Interval.Finset import Mathlib.Algebra.Order.Interval.Multiset @@ -2577,6 +2579,7 @@ import Mathlib.Data.Real.EReal import Mathlib.Data.Real.GoldenRatio import Mathlib.Data.Real.Hyperreal import Mathlib.Data.Real.Irrational +import Mathlib.Data.Real.IsNonarchimedean import Mathlib.Data.Real.Pi.Bounds import Mathlib.Data.Real.Pi.Irrational import Mathlib.Data.Real.Pi.Leibniz diff --git a/Mathlib/Algebra/Order/Hom/Normed.lean b/Mathlib/Algebra/Order/Hom/Normed.lean new file mode 100644 index 0000000000000..b44ee621b367a --- /dev/null +++ b/Mathlib/Algebra/Order/Hom/Normed.lean @@ -0,0 +1,75 @@ +/- +Copyright (c) 2024 Yakov Pechersky. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yakov Pechersky +-/ +import Mathlib.Algebra.Order.Hom.Basic +import Mathlib.Analysis.Normed.Group.Basic + +/-! +# Constructing (semi)normed groups from (semi)normed homs + +This file defines constructions that upgrade `(Comm)Group` to `(Semi)Normed(Comm)Group` +using a `Group(Semi)normClass` when the codomain is the reals. + +See `Mathlib.Algebra.Order.Hom.Ultra` for further upgrades to nonarchimedean normed groups. + +-/ + +variable {F α : Type*} [FunLike F α ℝ] + +/-- Constructs a `SeminormedGroup` structure from a `GroupSeminormClass` on a `Group`. -/ +-- See note [reducible non-instances] +@[to_additive "Constructs a `SeminormedAddGroup` structure from an `AddGroupSeminormClass` on an +`AddGroup`."] +abbrev GroupSeminormClass.toSeminormedGroup [Group α] [GroupSeminormClass F α ℝ] + (f : F) : SeminormedGroup α where + norm := f + dist x y := f (x / y) + dist_eq _ _ := rfl + dist_self _ := by simp + dist_comm x y := by simp only [← map_inv_eq_map f (x / y), inv_div] + dist_triangle x y z := by simpa using map_mul_le_add f (x / y) (y / z) + +@[to_additive] +lemma GroupSeminormClass.toSeminormedGroup_norm_eq [Group α] [GroupSeminormClass F α ℝ] + (f : F) (x : α) : @norm _ (GroupSeminormClass.toSeminormedGroup f).toNorm x = f x := rfl + +/-- Constructs a `SeminormedCommGroup` structure from a `GroupSeminormClass` on a `CommGroup`. -/ +-- See note [reducible non-instances] +@[to_additive "Constructs a `SeminormedAddCommGroup` structure from an `AddGroupSeminormClass` on an +`AddCommGroup`."] +abbrev GroupSeminormClass.toSeminormedCommGroup [CommGroup α] [GroupSeminormClass F α ℝ] + (f : F) : SeminormedCommGroup α where + __ := GroupSeminormClass.toSeminormedGroup f + __ : CommGroup α := inferInstance + +@[to_additive] +lemma GroupSeminormClass.toSeminormedCommGroup_norm_eq [CommGroup α] [GroupSeminormClass F α ℝ] + (f : F) (x : α) : @norm _ (GroupSeminormClass.toSeminormedCommGroup f).toNorm x = f x := rfl + +/-- Constructs a `NormedGroup` structure from a `GroupNormClass` on a `Group`. -/ +-- See note [reducible non-instances] +@[to_additive "Constructs a `NormedAddGroup` structure from an `AddGroupNormClass` on an +`AddGroup`."] +abbrev GroupNormClass.toNormedGroup [Group α] [GroupNormClass F α ℝ] + (f : F) : NormedGroup α where + __ := GroupSeminormClass.toSeminormedGroup f + eq_of_dist_eq_zero h := div_eq_one.mp (eq_one_of_map_eq_zero f h) + +@[to_additive] +lemma GroupNormClass.toNormedGroup_norm_eq [Group α] [GroupNormClass F α ℝ] + (f : F) (x : α) : @norm _ (GroupNormClass.toNormedGroup f).toNorm x = f x := rfl + +/-- Constructs a `NormedCommGroup` structure from a `GroupNormClass` on a `CommGroup`. -/ +-- See note [reducible non-instances] +@[to_additive "Constructs a `NormedAddCommGroup` structure from an `AddGroupNormClass` on an +`AddCommGroup`."] +abbrev GroupNormClass.toNormedCommGroup [CommGroup α] [GroupNormClass F α ℝ] + (f : F) : NormedCommGroup α where + __ := GroupNormClass.toNormedGroup f + __ : CommGroup α := inferInstance + +@[to_additive] +lemma GroupNormClass.toNormedCommGroup_norm_eq [CommGroup α] [GroupNormClass F α ℝ] + (f : F) (x : α) : @norm _ (GroupNormClass.toNormedCommGroup f).toNorm x = f x := rfl diff --git a/Mathlib/Algebra/Order/Hom/Ultra.lean b/Mathlib/Algebra/Order/Hom/Ultra.lean new file mode 100644 index 0000000000000..a0a03cd82334a --- /dev/null +++ b/Mathlib/Algebra/Order/Hom/Ultra.lean @@ -0,0 +1,32 @@ +/- +Copyright (c) 2024 Yakov Pechersky. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yakov Pechersky +-/ +import Mathlib.Algebra.Order.Hom.Normed +import Mathlib.Topology.MetricSpace.Ultra.Basic + +/-! +# Constructing nonarchimedean (ultrametric) normed groups from nonarchimedean normed homs + +This file defines constructions that upgrade `Add(Comm)Group` to `(Semi)NormedAdd(Comm)Group` +using an `AddGroup(Semi)normClass` when the codomain is the reals and the hom is nonarchimedean. + +## Implementation details + +The lemmas need to assume `[Dist α]` to be able to be stated, so they take an extra argument +that shows that this distance instance is propositionally equal to the one that comes from the +hom-based `AddGroupSeminormClass.toSeminormedAddGroup f` construction. To help at use site, +the argument is an autoparam that resolves by definitional equality when using these constructions. +-/ + +variable {F α : Type*} [FunLike F α ℝ] + +/-- Proves that when a `SeminormedAddGroup` structure is constructed from an +`AddGroupSeminormClass` that satifies `IsNonarchimedean`, the group has an `IsUltrametricDist`. -/ +lemma AddGroupSeminormClass.isUltrametricDist [AddGroup α] [AddGroupSeminormClass F α ℝ] + [inst : Dist α] {f : F} (hna : IsNonarchimedean f) + (hd : inst = (AddGroupSeminormClass.toSeminormedAddGroup f).toDist := by rfl) : + IsUltrametricDist α := + ⟨fun x y z ↦ by simpa only [hd, dist_eq_norm, AddGroupSeminormClass.toSeminormedAddGroup_norm_eq, + ← sub_add_sub_cancel x y z] using hna _ _⟩ diff --git a/Mathlib/Analysis/Normed/Group/Ultra.lean b/Mathlib/Analysis/Normed/Group/Ultra.lean index 8881f035276fb..44bfd92a9f0fe 100644 --- a/Mathlib/Analysis/Normed/Group/Ultra.lean +++ b/Mathlib/Analysis/Normed/Group/Ultra.lean @@ -237,6 +237,54 @@ lemma norm_prod_le_of_forall_le_of_nonneg {s : Finset ι} {f : ι → M} {C : lift C to NNReal using h_nonneg exact nnnorm_prod_le_of_forall_le hC +/-- +Given a function `f : ι → M` and a nonempty finite set `t ⊆ ι`, we can always find `i ∈ t` such that +`‖∏ j in t, f j‖ ≤ ‖f i‖`. +-/ +@[to_additive "Given a function `f : ι → M` and a nonempty finite set `t ⊆ ι`, we can always find +`i ∈ t` such that `‖∑ j in t, f j‖ ≤ ‖f i‖`."] +theorem exists_norm_finset_prod_le_of_nonempty {t : Finset ι} (ht : t.Nonempty) (f : ι → M) : + ∃ i ∈ t, ‖∏ j in t, f j‖ ≤ ‖f i‖ := + match t.exists_mem_eq_sup' ht (‖f ·‖) with + |⟨j, hj, hj'⟩ => ⟨j, hj, (ht.norm_prod_le_sup'_norm f).trans (le_of_eq hj')⟩ + +/-- +Given a function `f : ι → M` and a finite set `t ⊆ ι`, we can always find `i : ι`, belonging to `t` +if `t` is nonempty, such that `‖∏ j in t, f j‖ ≤ ‖f i‖`. +-/ +@[to_additive "Given a function `f : ι → M` and a finite set `t ⊆ ι`, we can always find `i : ι`, +belonging to `t` if `t` is nonempty, such that `‖∑ j in t, f j‖ ≤ ‖f i‖`."] +theorem exists_norm_finset_prod_le (t : Finset ι) [Nonempty ι] (f : ι → M) : + ∃ i : ι, (t.Nonempty → i ∈ t) ∧ ‖∏ j in t, f j‖ ≤ ‖f i‖ := by + rcases t.eq_empty_or_nonempty with rfl | ht + · simp + exact (fun ⟨i, h, h'⟩ => ⟨i, fun _ ↦ h, h'⟩) <| exists_norm_finset_prod_le_of_nonempty ht f + +/-- +Given a function `f : ι → M` and a multiset `t : Multiset ι`, we can always find `i : ι`, belonging +to `t` if `t` is nonempty, such that `‖(s.map f).prod‖ ≤ ‖f i‖`. +-/ +@[to_additive "Given a function `f : ι → M` and a multiset `t : Multiset ι`, we can always find +`i : ι`, belonging to `t` if `t` is nonempty, such that `‖(s.map f).sum‖ ≤ ‖f i‖`."] +theorem exists_norm_multiset_prod_le (s : Multiset ι) [Nonempty ι] {f : ι → M} : + ∃ i : ι, (s ≠ 0 → i ∈ s) ∧ ‖(s.map f).prod‖ ≤ ‖f i‖ := by + inhabit ι + induction s using Multiset.induction_on with + | empty => simp + | @cons a t hM => + obtain ⟨M, hMs, hM⟩ := hM + by_cases hMa : ‖f M‖ ≤ ‖f a‖ + · refine ⟨a, by simp, ?_⟩ + · rw [Multiset.map_cons, Multiset.prod_cons] + exact le_trans (norm_mul_le_max _ _) (max_le (le_refl _) (le_trans hM hMa)) + · rw [not_le] at hMa + rcases eq_or_ne t 0 with rfl|ht + · exact ⟨a, by simp, by simp⟩ + · refine ⟨M, ?_, ?_⟩ + · simp [hMs ht] + rw [Multiset.map_cons, Multiset.prod_cons] + exact le_trans (norm_mul_le_max _ _) (max_le hMa.le hM) + @[to_additive] lemma norm_tprod_le (f : ι → M) : ‖∏' i, f i‖ ≤ ⨆ i, ‖f i‖ := by rcases isEmpty_or_nonempty ι with hι | hι diff --git a/Mathlib/Data/Real/IsNonarchimedean.lean b/Mathlib/Data/Real/IsNonarchimedean.lean new file mode 100644 index 0000000000000..ca4e42c432aab --- /dev/null +++ b/Mathlib/Data/Real/IsNonarchimedean.lean @@ -0,0 +1,114 @@ +/- +Copyright (c) 2024 María Inés de Frutos-Fernández. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: María Inés de Frutos-Fernández +-/ +import Mathlib.Algebra.Order.Hom.Ultra +import Mathlib.Analysis.Normed.Group.Ultra + +/-! +# Nonarchimedean functions + +A function `f : R → ℝ≥0` is nonarchimedean if it satisfies the strong triangle inequality +`f (r + s) ≤ max (f r) (f s)` for all `r s : R`. This file proves basic properties of +nonarchimedean functions. +-/ + +namespace IsNonarchimedean + +open IsUltrametricDist + +/-- A nonarchimedean function satisfies the triangle inequality. -/ +theorem add_le {α : Type*} [Add α] {f : α → ℝ} (hf : ∀ x : α, 0 ≤ f x) + (hna : IsNonarchimedean f) {a b : α} : f (a + b) ≤ f a + f b := by + apply le_trans (hna _ _) + rw [max_le_iff, le_add_iff_nonneg_right, le_add_iff_nonneg_left] + exact ⟨hf _, hf _⟩ + +/-- If `f` is a nonarchimedean additive group seminorm on `α`, then for every `n : ℕ` and `a : α`, + we have `f (n • a) ≤ (f a)`. -/ +theorem nsmul_le {F α : Type*} [AddGroup α] [FunLike F α ℝ] + [AddGroupSeminormClass F α ℝ] {f : F} (hna : IsNonarchimedean f) {n : ℕ} {a : α} : + f (n • a) ≤ f a := by + let _ := AddGroupSeminormClass.toSeminormedAddGroup f + have := AddGroupSeminormClass.isUltrametricDist hna + simp only [← AddGroupSeminormClass.toSeminormedAddGroup_norm_eq] + exact norm_nsmul_le _ _ + +/-- If `f` is a nonarchimedean additive group seminorm on `α`, then for every `n : ℕ` and `a : α`, + we have `f (n * a) ≤ (f a)`. -/ +theorem nmul_le {F α : Type*} [Ring α] [FunLike F α ℝ] [AddGroupSeminormClass F α ℝ] + {f : F} (hna : IsNonarchimedean f) {n : ℕ} {a : α} : f (n * a) ≤ f a := by + rw [← nsmul_eq_mul] + exact nsmul_le hna + +/-- If `f` is a nonarchimedean additive group seminorm on `α` and `x y : α` are such that + `f x ≠ f y`, then `f (x + y) = max (f x) (f y)`. -/ +theorem add_eq_max_of_ne {F α : Type*} [AddGroup α] [FunLike F α ℝ] + [AddGroupSeminormClass F α ℝ] {f : F} (hna : IsNonarchimedean f) {x y : α} (hne : f x ≠ f y) : + f (x + y) = max (f x) (f y) := by + let _ := AddGroupSeminormClass.toSeminormedAddGroup f + have := AddGroupSeminormClass.isUltrametricDist hna + simp only [← AddGroupSeminormClass.toSeminormedAddGroup_norm_eq] at hne ⊢ + exact norm_add_eq_max_of_norm_ne_norm hne + +/-- Given a nonarchimedean additive group seminorm `f` on `α`, a function `g : β → α` and a finset + `t : Finset β`, we can always find `b : β`, belonging to `t` if `t` is nonempty, such that + `f (t.sum g) ≤ f (g b)` . -/ +theorem finset_image_add {F α β : Type*} [AddCommGroup α] [FunLike F α ℝ] + [AddGroupSeminormClass F α ℝ] [Nonempty β] {f : F} (hna : IsNonarchimedean f) + (g : β → α) (t : Finset β) : + ∃ b : β, (t.Nonempty → b ∈ t) ∧ f (t.sum g) ≤ f (g b) := by + let _ := AddGroupSeminormClass.toSeminormedAddCommGroup f + have := AddGroupSeminormClass.isUltrametricDist hna + simp only [← AddGroupSeminormClass.toSeminormedAddCommGroup_norm_eq] + apply exists_norm_finset_sum_le + +/-- Given a nonarchimedean additive group seminorm `f` on `α`, a function `g : β → α` and a + nonempty finset `t : Finset β`, we can always find `b : β` belonging to `t` such that + `f (t.sum g) ≤ f (g b)` . -/ +theorem finset_image_add_of_nonempty {F α β : Type*} [AddCommGroup α] [FunLike F α ℝ] + [AddGroupSeminormClass F α ℝ] [Nonempty β] {f : F} (hna : IsNonarchimedean f) + (g : β → α) {t : Finset β} (ht : t.Nonempty) : + ∃ b : β, (b ∈ t) ∧ f (t.sum g) ≤ f (g b) := by + obtain ⟨b, hbt, hbf⟩ := finset_image_add hna g t + exact ⟨b, hbt ht, hbf⟩ + +/-- Given a nonarchimedean additive group seminorm `f` on `α`, a function `g : β → α` and a + multiset `s : Multiset β`, we can always find `b : β`, belonging to `s` if `s` is nonempty, + such that `f (t.sum g) ≤ f (g b)` . -/ +theorem multiset_image_add {F α β : Type*} [AddCommGroup α] [FunLike F α ℝ] + [AddGroupSeminormClass F α ℝ] [Nonempty β] {f : F} (hna : IsNonarchimedean f) + (g : β → α) (s : Multiset β) : + ∃ b : β, (s ≠ 0 → b ∈ s) ∧ f (Multiset.map g s).sum ≤ f (g b) := by + let _ := AddGroupSeminormClass.toSeminormedAddCommGroup f + have := AddGroupSeminormClass.isUltrametricDist hna + simp only [← AddGroupSeminormClass.toSeminormedAddCommGroup_norm_eq] + apply exists_norm_multiset_sum_le + +/-- Given a nonarchimedean additive group seminorm `f` on `α`, a function `g : β → α` and a + nonempty multiset `s : Multiset β`, we can always find `b : β` belonging to `s` such that + `f (t.sum g) ≤ f (g b)` . -/ +theorem multiset_image_add_of_nonempty {F α β : Type*} [AddCommGroup α] [FunLike F α ℝ] + [AddGroupSeminormClass F α ℝ] [Nonempty β] {f : F} (hna : IsNonarchimedean f) + (g : β → α) {s : Multiset β} (hs : s ≠ 0) : + ∃ b : β, (b ∈ s) ∧ f (Multiset.map g s).sum ≤ f (g b) := by + obtain ⟨b, hbs, hbf⟩ := multiset_image_add hna g s + exact ⟨b, hbs hs, hbf⟩ + +/-- If `f` is a nonarchimedean additive group seminorm on a commutative ring `α`, `n : ℕ`, and + `a b : α`, then we can find `m : ℕ` such that `m ≤ n` and + `f ((a + b) ^ n) ≤ (f (a ^ m)) * (f (b ^ (n - m)))`. -/ +theorem add_pow_le {F α : Type*} [CommRing α] [FunLike F α ℝ] + [RingSeminormClass F α ℝ] {f : F} (hna : IsNonarchimedean f) (n : ℕ) (a b : α) : + ∃ m < n + 1, f ((a + b) ^ n) ≤ f (a ^ m) * f (b ^ (n - m)) := by + obtain ⟨m, hm_lt, hM⟩ := finset_image_add hna + (fun m => a ^ m * b ^ (n - m) * ↑(n.choose m)) (Finset.range (n + 1)) + simp only [Finset.nonempty_range_iff, ne_eq, Nat.succ_ne_zero, not_false_iff, Finset.mem_range, + if_true, forall_true_left] at hm_lt + refine ⟨m, hm_lt, ?_⟩ + simp only [← add_pow] at hM + rw [mul_comm] at hM + exact le_trans hM (le_trans (nmul_le hna) (map_mul_le_mul _ _ _)) + +end IsNonarchimedean From a60ffdad9dc7fd9d6124ce14b549cb27add5cad2 Mon Sep 17 00:00:00 2001 From: Pietro Monticone <38562595+pitmonticone@users.noreply.github.com> Date: Mon, 21 Oct 2024 14:44:12 +0000 Subject: [PATCH 239/505] chore: add dependabot (#18013) This PR introduces a Dependabot configuration specifically for managing dependencies in GitHub Actions workflows. It automates the process of checking for updates to GitHub Actions dependencies. --- .github/dependabot.yml | 9 +++++++++ 1 file changed, 9 insertions(+) create mode 100644 .github/dependabot.yml diff --git a/.github/dependabot.yml b/.github/dependabot.yml new file mode 100644 index 0000000000000..4e43547be01ea --- /dev/null +++ b/.github/dependabot.yml @@ -0,0 +1,9 @@ +version: 2 # Specifies the version of the Dependabot configuration file format + +updates: + # Configuration for dependency updates + - package-ecosystem: "github-actions" # Specifies the ecosystem to check for updates + directory: "/" # Specifies the directory to check for dependencies; "/" means the root directory + schedule: + # Check for updates to GitHub Actions every month + interval: "monthly" From 69a4014a0c0066ae5ec391a0ac8f6865a1a12a28 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Mon, 21 Oct 2024 14:58:55 +0000 Subject: [PATCH 240/505] chore(CategoryTheory/Adjunction): homEquiv_unit is no longer a global simp lemma (#17761) This PR removes the global `simp` attribute to lemmas `homEquiv_unit` and `homEquiv_counit`. These lemmas express the bijection `homEquiv` in terms of the unit and counit of the adjunction. In situations where adjunctions are defined using the constructor `mkOfHomEquiv`, we do not want these lemmas as global simp. However, when we study adjunctions defined using units and counits, or when proving general API results about them, they can be made local simp. --- Mathlib/CategoryTheory/Adjunction/Basic.lean | 23 ++++++++++++------- Mathlib/CategoryTheory/Adjunction/Comma.lean | 2 ++ .../Adjunction/FullyFaithful.lean | 2 ++ .../Adjunction/Lifting/Left.lean | 2 ++ .../Adjunction/Lifting/Right.lean | 1 + Mathlib/CategoryTheory/Adjunction/Limits.lean | 2 ++ .../CategoryTheory/Adjunction/Opposites.lean | 2 ++ .../CategoryTheory/Adjunction/Reflective.lean | 2 +- .../CategoryTheory/Adjunction/Restrict.lean | 2 ++ Mathlib/CategoryTheory/Adjunction/Unique.lean | 2 ++ Mathlib/CategoryTheory/Closed/Ideal.lean | 5 ++-- Mathlib/CategoryTheory/Limits/Final.lean | 2 ++ Mathlib/CategoryTheory/Limits/HasLimits.lean | 2 +- .../Localization/Bousfield.lean | 1 + .../CategoryTheory/Monad/Comonadicity.lean | 6 ++--- Mathlib/CategoryTheory/Monad/Monadicity.lean | 4 ++-- Mathlib/CategoryTheory/Monoidal/Functor.lean | 2 ++ .../Monoidal/NaturalTransformation.lean | 4 +++- .../CategoryTheory/Preadditive/Injective.lean | 3 ++- .../Sites/PreservesSheafification.lean | 8 +++---- Mathlib/Topology/Sheaves/Presheaf.lean | 4 ++-- 21 files changed, 56 insertions(+), 25 deletions(-) diff --git a/Mathlib/CategoryTheory/Adjunction/Basic.lean b/Mathlib/CategoryTheory/Adjunction/Basic.lean index 3d9c1506b6275..d7c1293793302 100644 --- a/Mathlib/CategoryTheory/Adjunction/Basic.lean +++ b/Mathlib/CategoryTheory/Adjunction/Basic.lean @@ -151,7 +151,7 @@ namespace Adjunction attribute [reassoc (attr := simp)] left_triangle_components right_triangle_components /-- The hom set equivalence associated to an adjunction. -/ -@[simps] +@[simps (config := .lemmasOnly)] def homEquiv {F : C ⥤ D} {G : D ⥤ C} (adj : F ⊣ G) (X : C) (Y : D) : (F.obj X ⟶ Y) ≃ (X ⟶ G.obj Y) where toFun := fun f => adj.unit.app X ≫ G.map f @@ -165,7 +165,17 @@ def homEquiv {F : C ⥤ D} {G : D ⥤ C} (adj : F ⊣ G) (X : C) (Y : D) : alias homEquiv_unit := homEquiv_apply alias homEquiv_counit := homEquiv_symm_apply -attribute [simp] homEquiv_unit homEquiv_counit + +end Adjunction + +-- These lemmas are not global simp lemmas because certain adjunctions +-- are constructed using `Adjunction.mkOfHomEquiv`, and we certainly +-- do not want `dsimp` to apply `homEquiv_unit` or `homEquiv_counit` +-- in that case. However, when proving general API results about adjunctions, +-- it may be advisable to add a local simp attribute to these lemmas. +attribute [local simp] Adjunction.homEquiv_unit Adjunction.homEquiv_counit + +namespace Adjunction section @@ -289,12 +299,6 @@ structure CoreHomEquivUnitCounit (F : C ⥤ D) (G : D ⥤ C) where /-- The relationship between the counit and hom set equivalence of an adjunction -/ homEquiv_counit : ∀ {X Y g}, (homEquiv X Y).symm g = F.map g ≫ counit.app Y := by aesop_cat -namespace CoreHomEquivUnitCounit - -attribute [simp] homEquiv_unit homEquiv_counit - -end CoreHomEquivUnitCounit - /-- This is an auxiliary data structure useful for constructing adjunctions. See `Adjunction.mkOfHomEquiv`. This structure won't typically be used anywhere else. @@ -363,6 +367,8 @@ end CoreUnitCounit variable {F : C ⥤ D} {G : D ⥤ C} +attribute [local simp] CoreHomEquivUnitCounit.homEquiv_unit CoreHomEquivUnitCounit.homEquiv_counit + /-- Construct an adjunction from the data of a `CoreHomEquivUnitCounit`, i.e. a hom set equivalence, unit and counit natural transformations together with proofs of the equalities @@ -401,6 +407,7 @@ def mkOfHomEquiv (adj : CoreHomEquiv F G) : F ⊣ G := homEquiv_unit := fun {X Y f} => by simp [← adj.homEquiv_naturality_right] homEquiv_counit := fun {X Y f} => by simp [← adj.homEquiv_naturality_left_symm] } +@[simp] lemma mkOfHomEquiv_homEquiv (adj : CoreHomEquiv F G) : (mkOfHomEquiv adj).homEquiv = adj.homEquiv := by ext X Y g diff --git a/Mathlib/CategoryTheory/Adjunction/Comma.lean b/Mathlib/CategoryTheory/Adjunction/Comma.lean index b74e9b326b43d..3b59a2dd44bdc 100644 --- a/Mathlib/CategoryTheory/Adjunction/Comma.lean +++ b/Mathlib/CategoryTheory/Adjunction/Comma.lean @@ -125,6 +125,8 @@ section variable {F : C ⥤ D} +attribute [local simp] Adjunction.homEquiv_unit Adjunction.homEquiv_counit + /-- Given a left adjoint to `G`, we can construct an initial object in each structured arrow category on `G`. -/ def mkInitialOfLeftAdjoint (h : F ⊣ G) (A : C) : diff --git a/Mathlib/CategoryTheory/Adjunction/FullyFaithful.lean b/Mathlib/CategoryTheory/Adjunction/FullyFaithful.lean index 0c46fa9aba37c..7c789d7c00fcf 100644 --- a/Mathlib/CategoryTheory/Adjunction/FullyFaithful.lean +++ b/Mathlib/CategoryTheory/Adjunction/FullyFaithful.lean @@ -40,6 +40,8 @@ open Category open Opposite +attribute [local simp] Adjunction.homEquiv_unit Adjunction.homEquiv_counit + variable {C : Type u₁} [Category.{v₁} C] variable {D : Type u₂} [Category.{v₂} D] variable {L : C ⥤ D} {R : D ⥤ C} (h : L ⊣ R) diff --git a/Mathlib/CategoryTheory/Adjunction/Lifting/Left.lean b/Mathlib/CategoryTheory/Adjunction/Lifting/Left.lean index e8cd5e6498b69..f6b0b6a9aedb8 100644 --- a/Mathlib/CategoryTheory/Adjunction/Lifting/Left.lean +++ b/Mathlib/CategoryTheory/Adjunction/Lifting/Left.lean @@ -157,6 +157,8 @@ noncomputable def constructLeftAdjointEquiv [∀ X : B, RegularEpi (adj₁.couni apply eq_comm _ ≃ (X ⟶ R.obj Y) := (Cofork.IsColimit.homIso (counitCoequalises adj₁ X) _).symm +attribute [local simp] Adjunction.homEquiv_counit + /-- Construct the left adjoint to `R`, with object map `constructLeftAdjointObj`. -/ noncomputable def constructLeftAdjoint [∀ X : B, RegularEpi (adj₁.counit.app X)] : B ⥤ A := by refine Adjunction.leftAdjointOfEquiv (fun X Y => constructLeftAdjointEquiv R _ adj₁ adj₂ Y X) ?_ diff --git a/Mathlib/CategoryTheory/Adjunction/Lifting/Right.lean b/Mathlib/CategoryTheory/Adjunction/Lifting/Right.lean index 3d3893abeac50..83076882dba44 100644 --- a/Mathlib/CategoryTheory/Adjunction/Lifting/Right.lean +++ b/Mathlib/CategoryTheory/Adjunction/Lifting/Right.lean @@ -152,6 +152,7 @@ noncomputable def constructRightAdjoint [∀ X : B, RegularMono (adj₁.unit.app rw [constructRightAdjointEquiv_symm_apply, constructRightAdjointEquiv_symm_apply, Equiv.symm_apply_eq, Subtype.ext_iff] dsimp + simp only [Adjunction.homEquiv_unit, Adjunction.homEquiv_counit] erw [Fork.IsLimit.homIso_natural, Fork.IsLimit.homIso_natural] simp only [Fork.ofι_pt, Functor.map_comp, assoc, limit.cone_x] erw [adj₂.homEquiv_naturality_left, Equiv.rightInverse_symm] diff --git a/Mathlib/CategoryTheory/Adjunction/Limits.lean b/Mathlib/CategoryTheory/Adjunction/Limits.lean index 6a2f33e8ce6ad..f57ec0e073cee 100644 --- a/Mathlib/CategoryTheory/Adjunction/Limits.lean +++ b/Mathlib/CategoryTheory/Adjunction/Limits.lean @@ -296,6 +296,8 @@ end ArbitraryUniverse variable {C : Type u₁} [Category.{v₀} C] {D : Type u₂} [Category.{v₀} D] {F : C ⥤ D} {G : D ⥤ C} (adj : F ⊣ G) +attribute [local simp] homEquiv_unit homEquiv_counit + -- Note: this is natural in K, but we do not yet have the tools to formulate that. /-- When `F ⊣ G`, the functor associating to each `Y` the cocones over `K ⋙ F` with cone point `Y` diff --git a/Mathlib/CategoryTheory/Adjunction/Opposites.lean b/Mathlib/CategoryTheory/Adjunction/Opposites.lean index f4362585c27a4..615dd8ae5d899 100644 --- a/Mathlib/CategoryTheory/Adjunction/Opposites.lean +++ b/Mathlib/CategoryTheory/Adjunction/Opposites.lean @@ -27,6 +27,8 @@ variable {C : Type u₁} [Category.{v₁} C] {D : Type u₂} [Category.{v₂} D] namespace CategoryTheory.Adjunction +attribute [local simp] homEquiv_unit homEquiv_counit + /-- If `G.op` is adjoint to `F.op` then `F` is adjoint to `G`. -/ @[simps! unit_app counit_app] def adjointOfOpAdjointOp (F : C ⥤ D) (G : D ⥤ C) (h : G.op ⊣ F.op) : F ⊣ G := diff --git a/Mathlib/CategoryTheory/Adjunction/Reflective.lean b/Mathlib/CategoryTheory/Adjunction/Reflective.lean index 2ddd0414b0f2f..1c0c1c350dd0c 100644 --- a/Mathlib/CategoryTheory/Adjunction/Reflective.lean +++ b/Mathlib/CategoryTheory/Adjunction/Reflective.lean @@ -113,7 +113,7 @@ def unitCompPartialBijectiveAux [Reflective i] (A : C) (B : D) : theorem unitCompPartialBijectiveAux_symm_apply [Reflective i] {A : C} {B : D} (f : i.obj ((reflector i).obj A) ⟶ i.obj B) : (unitCompPartialBijectiveAux _ _).symm f = (reflectorAdjunction i).unit.app A ≫ f := by - simp [unitCompPartialBijectiveAux] + simp [unitCompPartialBijectiveAux, Adjunction.homEquiv_unit] /-- If `i` has a reflector `L`, then the function `(i.obj (L.obj A) ⟶ B) → (A ⟶ B)` given by precomposing with `η.app A` is a bijection provided `B` is in the essential image of `i`. diff --git a/Mathlib/CategoryTheory/Adjunction/Restrict.lean b/Mathlib/CategoryTheory/Adjunction/Restrict.lean index d3e8258ab5062..af50916ad19c8 100644 --- a/Mathlib/CategoryTheory/Adjunction/Restrict.lean +++ b/Mathlib/CategoryTheory/Adjunction/Restrict.lean @@ -27,6 +27,8 @@ variable {iC : C ⥤ C'} {iD : D ⥤ D'} {L' : C' ⥤ D'} {R' : D' ⥤ C'} (adj : L' ⊣ R') (hiC : iC.FullyFaithful) (hiD : iD.FullyFaithful) {L : C ⥤ D} {R : D ⥤ C} (comm1 : iC ⋙ L' ≅ L ⋙ iD) (comm2 : iD ⋙ R' ≅ R ⋙ iC) +attribute [local simp] homEquiv_unit homEquiv_counit + /-- If `C` is a full subcategory of `C'` and `D` is a full subcategory of `D'`, then we can restrict an adjunction `L' ⊣ R'` where `L' : C' ⥤ D'` and `R' : D' ⥤ C'` to `C` and `D`. The construction here is slightly more general, in that `C` is required only to have a full and diff --git a/Mathlib/CategoryTheory/Adjunction/Unique.lean b/Mathlib/CategoryTheory/Adjunction/Unique.lean index 9c4b07886fc7d..7b1a98e824fa1 100644 --- a/Mathlib/CategoryTheory/Adjunction/Unique.lean +++ b/Mathlib/CategoryTheory/Adjunction/Unique.lean @@ -26,6 +26,8 @@ variable {C D : Type*} [Category C] [Category D] namespace CategoryTheory.Adjunction +attribute [local simp] homEquiv_unit homEquiv_counit + /-- 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' := ((conjugateIsoEquiv adj1 adj2).symm (Iso.refl G)).symm diff --git a/Mathlib/CategoryTheory/Closed/Ideal.lean b/Mathlib/CategoryTheory/Closed/Ideal.lean index 588c491499a4f..aea4dfe3a8e28 100644 --- a/Mathlib/CategoryTheory/Closed/Ideal.lean +++ b/Mathlib/CategoryTheory/Closed/Ideal.lean @@ -202,7 +202,7 @@ theorem bijection_symm_apply_id (A B : C) : -- Porting note: added dsimp only [Functor.comp_obj] rw [prod.comp_lift_assoc, prod.lift_snd, prod.lift_fst_assoc, prod.lift_fst_comp_snd_comp, - ← Adjunction.eq_unit_comp_map_iff, Iso.comp_inv_eq, assoc] + Adjunction.homEquiv_counit, ← Adjunction.eq_unit_comp_map_iff, Iso.comp_inv_eq, assoc] rw [PreservesLimitPair.iso_hom i ((reflector i).obj A) ((reflector i).obj B)] apply Limits.prod.hom_ext · rw [Limits.prod.map_fst, assoc, assoc, prodComparison_fst, ← i.map_comp, prodComparison_fst] @@ -217,7 +217,8 @@ theorem bijection_natural (A B : C) (X X' : D) (f : (reflector i).obj (A ⨯ B) erw [homEquiv_symm_apply_eq, homEquiv_symm_apply_eq, homEquiv_apply_eq, homEquiv_apply_eq, homEquiv_symm_apply_eq, homEquiv_symm_apply_eq, homEquiv_apply_eq, homEquiv_apply_eq] apply i.map_injective - rw [Functor.FullyFaithful.map_preimage, i.map_comp] + rw [Functor.FullyFaithful.map_preimage, i.map_comp, + Adjunction.homEquiv_unit, Adjunction.homEquiv_unit] simp only [comp_id, Functor.map_comp, Functor.FullyFaithful.map_preimage, assoc] rw [← assoc, ← assoc, curry_natural_right _ (i.map g), unitCompPartialBijective_natural, uncurry_natural_right, ← assoc, curry_natural_right, diff --git a/Mathlib/CategoryTheory/Limits/Final.lean b/Mathlib/CategoryTheory/Limits/Final.lean index 965ee9e0e669e..12c39f6153dc8 100644 --- a/Mathlib/CategoryTheory/Limits/Final.lean +++ b/Mathlib/CategoryTheory/Limits/Final.lean @@ -116,6 +116,8 @@ theorem initial_of_final_op (F : C ⥤ D) [Final F.op] : Initial F := @isConnected_of_isConnected_op _ _ (isConnected_of_equivalent (costructuredArrowOpEquivalence F d).symm) } +attribute [local simp] Adjunction.homEquiv_unit Adjunction.homEquiv_counit + /-- If a functor `R : D ⥤ C` is a right adjoint, it is final. -/ theorem final_of_adjunction {L : C ⥤ D} {R : D ⥤ C} (adj : L ⊣ R) : Final R := { out := fun c => diff --git a/Mathlib/CategoryTheory/Limits/HasLimits.lean b/Mathlib/CategoryTheory/Limits/HasLimits.lean index 75d8b50abec23..06940a016159d 100644 --- a/Mathlib/CategoryTheory/Limits/HasLimits.lean +++ b/Mathlib/CategoryTheory/Limits/HasLimits.lean @@ -539,7 +539,7 @@ def isLimitConeOfAdj (F : J ⥤ C) : have eq := NatTrans.congr_app (adj.counit.naturality s.π) j have eq' := NatTrans.congr_app (adj.left_triangle_components s.pt) j dsimp at eq eq' ⊢ - rw [assoc, eq, reassoc_of% eq'] + rw [adj.homEquiv_unit, assoc, eq, reassoc_of% eq'] uniq s m hm := (adj.homEquiv _ _).symm.injective (by ext j; simpa using hm j) end Adjunction diff --git a/Mathlib/CategoryTheory/Localization/Bousfield.lean b/Mathlib/CategoryTheory/Localization/Bousfield.lean index 1dd1de3c1799e..209e43b314608 100644 --- a/Mathlib/CategoryTheory/Localization/Bousfield.lean +++ b/Mathlib/CategoryTheory/Localization/Bousfield.lean @@ -112,6 +112,7 @@ lemma W_adj_unit_app (X : D) : W (· ∈ Set.range F.obj) (adj.unit.app X) := by rintro _ ⟨Y, rfl⟩ convert ((Functor.FullyFaithful.ofFullyFaithful F).homEquiv.symm.trans (adj.homEquiv X Y)).bijective using 1 + dsimp [Adjunction.homEquiv] aesop lemma W_iff_isIso_map {X Y : D} (f : X ⟶ Y) : diff --git a/Mathlib/CategoryTheory/Monad/Comonadicity.lean b/Mathlib/CategoryTheory/Monad/Comonadicity.lean index 7aeb43cea5361..2649f2cfb69b3 100644 --- a/Mathlib/CategoryTheory/Monad/Comonadicity.lean +++ b/Mathlib/CategoryTheory/Monad/Comonadicity.lean @@ -117,7 +117,7 @@ def rightAdjointComparison · apply comparisonRightAdjointHomEquiv · intro A B B' g h apply equalizer.hom_ext - simp + simp [Adjunction.homEquiv_unit] /-- Provided we have the appropriate equalizers, we have an adjunction to the comparison functor. -/ @@ -161,7 +161,7 @@ theorem comparisonAdjunction_counit_f (adj.unit.app (G.obj A.A))] (A : adj.toComonad.Coalgebra) : ((comparisonAdjunction adj).counit.app A).f = (beckEqualizer A).lift (counitFork A) := by - simp + simp [Adjunction.homEquiv_counit] variable (adj) @@ -205,7 +205,7 @@ theorem comparisonAdjunction_unit_app change equalizer.lift ((adj.homEquiv B _) (𝟙 _)) _ ≫ equalizer.ι _ _ = equalizer.lift _ _ ≫ equalizer.ι _ _ - simp + simp [Adjunction.homEquiv_unit] end ComonadicityInternal diff --git a/Mathlib/CategoryTheory/Monad/Monadicity.lean b/Mathlib/CategoryTheory/Monad/Monadicity.lean index 134144358f808..329b6f0907bb9 100644 --- a/Mathlib/CategoryTheory/Monad/Monadicity.lean +++ b/Mathlib/CategoryTheory/Monad/Monadicity.lean @@ -130,7 +130,7 @@ def leftAdjointComparison -- `Category.assoc`. -- dsimp [comparisonLeftAdjointHomEquiv] -- rw [← adj.homEquiv_naturality_right, Category.assoc] - simp [Cofork.IsColimit.homIso] + simp [Cofork.IsColimit.homIso, Adjunction.homEquiv_unit] /-- Provided we have the appropriate coequalizers, we have an adjunction to the comparison functor. -/ @@ -225,7 +225,7 @@ theorem comparisonAdjunction_counit_app change coequalizer.π _ _ ≫ coequalizer.desc ((adj.homEquiv _ B).symm (𝟙 _)) _ = coequalizer.π _ _ ≫ coequalizer.desc _ _ - simp + simp [Adjunction.homEquiv_counit] end MonadicityInternal diff --git a/Mathlib/CategoryTheory/Monoidal/Functor.lean b/Mathlib/CategoryTheory/Monoidal/Functor.lean index 11f402b27a233..4f367c8c6a40a 100644 --- a/Mathlib/CategoryTheory/Monoidal/Functor.lean +++ b/Mathlib/CategoryTheory/Monoidal/Functor.lean @@ -736,10 +736,12 @@ noncomputable def monoidalAdjoint : instance [F.IsEquivalence] : IsIso (monoidalAdjoint F h).ε := by dsimp + rw [Adjunction.homEquiv_unit] infer_instance instance (X Y : D) [F.IsEquivalence] : IsIso ((monoidalAdjoint F h).μ X Y) := by dsimp + rw [Adjunction.homEquiv_unit] infer_instance /-- If a monoidal functor `F` is an equivalence of categories then its inverse is also monoidal. -/ diff --git a/Mathlib/CategoryTheory/Monoidal/NaturalTransformation.lean b/Mathlib/CategoryTheory/Monoidal/NaturalTransformation.lean index d41970b459890..eb01593defb37 100644 --- a/Mathlib/CategoryTheory/Monoidal/NaturalTransformation.lean +++ b/Mathlib/CategoryTheory/Monoidal/NaturalTransformation.lean @@ -171,6 +171,7 @@ def monoidalUnit : F.map_tensor, IsIso.hom_inv_id_assoc, ← tensor_comp_assoc, Adjunction.left_triangle_components, tensorHom_id, id_whiskerRight, IsIso.inv_hom_id, map_id] + unit := by simp [Adjunction.homEquiv_unit] /-- The unit of a adjunction can be upgraded to a monoidal natural transformation. -/ @[simps] @@ -188,7 +189,8 @@ def monoidalCounit : unit := by have eq := h.counit.naturality F.ε dsimp at eq ⊢ - rw [map_inv, map_comp, assoc, assoc, map_inv, ← cancel_mono F.ε, assoc, assoc, assoc, ← eq, + rw [Adjunction.homEquiv_unit, map_inv, map_comp, assoc, assoc, map_inv, + ← cancel_mono F.ε, assoc, assoc, assoc, ← eq, IsIso.inv_hom_id_assoc, Adjunction.left_triangle_components, comp_id, id_comp] instance [F.IsEquivalence] : IsIso (monoidalUnit F h) := by diff --git a/Mathlib/CategoryTheory/Preadditive/Injective.lean b/Mathlib/CategoryTheory/Preadditive/Injective.lean index 3b33277d8bf28..4ee0d211f19f3 100644 --- a/Mathlib/CategoryTheory/Preadditive/Injective.lean +++ b/Mathlib/CategoryTheory/Preadditive/Injective.lean @@ -187,7 +187,8 @@ variable {L : C ⥤ D} {R : D ⥤ C} [PreservesMonomorphisms L] theorem injective_of_adjoint (adj : L ⊣ R) (J : D) [Injective J] : Injective <| R.obj J := ⟨fun {A} {_} g f im => ⟨adj.homEquiv _ _ (factorThru ((adj.homEquiv A J).symm g) (L.map f)), - (adj.homEquiv _ _).symm.injective (by simp)⟩⟩ + (adj.homEquiv _ _).symm.injective + (by simp [Adjunction.homEquiv_unit, Adjunction.homEquiv_counit])⟩⟩ end Adjunction diff --git a/Mathlib/CategoryTheory/Sites/PreservesSheafification.lean b/Mathlib/CategoryTheory/Sites/PreservesSheafification.lean index 7e02c805dd32b..010ddc167e260 100644 --- a/Mathlib/CategoryTheory/Sites/PreservesSheafification.lean +++ b/Mathlib/CategoryTheory/Sites/PreservesSheafification.lean @@ -170,7 +170,8 @@ lemma sheafComposeNatTrans_fac (P : Cᵒᵖ ⥤ A) : adj₂.unit.app (P ⋙ F) ≫ (sheafToPresheaf J B).map ((sheafComposeNatTrans J F adj₁ adj₂).app P) = whiskerRight (adj₁.unit.app P) F := by - simp [sheafComposeNatTrans, -sheafToPresheaf_obj, -sheafToPresheaf_map] + simp [sheafComposeNatTrans, -sheafToPresheaf_obj, -sheafToPresheaf_map, + Adjunction.homEquiv_counit] lemma sheafComposeNatTrans_app_uniq (P : Cᵒᵖ ⥤ A) (α : G₂.obj (P ⋙ F) ⟶ (sheafCompose J F).obj (G₁.obj P)) @@ -264,9 +265,8 @@ lemma sheafToPresheaf_map_sheafComposeNatTrans_eq_sheafifyCompIso_inv (P : Cᵒ rfl apply ((plusPlusAdjunction J E).homEquiv _ _).injective convert sheafComposeNatTrans_fac J F (plusPlusAdjunction J D) (plusPlusAdjunction J E) P - all_goals - dsimp [plusPlusAdjunction] - simp + dsimp [plusPlusAdjunction] + simp instance (P : Cᵒᵖ ⥤ D) : IsIso ((sheafComposeNatTrans J F (plusPlusAdjunction J D) (plusPlusAdjunction J E)).app P) := by diff --git a/Mathlib/Topology/Sheaves/Presheaf.lean b/Mathlib/Topology/Sheaves/Presheaf.lean index 6b2b727509484..b4a85d939a41c 100644 --- a/Mathlib/Topology/Sheaves/Presheaf.lean +++ b/Mathlib/Topology/Sheaves/Presheaf.lean @@ -243,7 +243,7 @@ theorem toPushforwardOfIso_app {X Y : TopCat} (H₁ : X ≅ Y) {ℱ : X.Presheaf ℱ.map (eqToHom (by simp [Opens.map, Set.preimage_preimage])) ≫ H₂.app (op ((Opens.map H₁.inv).obj (unop U))) := by delta toPushforwardOfIso - simp [-Functor.map_comp, ← Functor.map_comp_assoc] + simp [-Functor.map_comp, ← Functor.map_comp_assoc, Adjunction.homEquiv_unit] rfl /-- If `H : X ≅ Y` is a homeomorphism, @@ -259,7 +259,7 @@ theorem pushforwardToOfIso_app {X Y : TopCat} (H₁ : X ≅ Y) {ℱ : Y.Presheaf (pushforwardToOfIso H₁ H₂).app U = H₂.app (op ((Opens.map H₁.inv).obj (unop U))) ≫ 𝒢.map (eqToHom (by simp [Opens.map, Set.preimage_preimage])) := by - simp [pushforwardToOfIso, Equivalence.toAdjunction] + simp [pushforwardToOfIso, Equivalence.toAdjunction, Adjunction.homEquiv_counit] end Iso From 7004cd3fee7719c0bda62a4cff6ed52b4b7e0036 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Mon, 21 Oct 2024 15:16:21 +0000 Subject: [PATCH 241/505] feat: add missing instances for `ContinuousMap` and `BoundedContinuousFunction` (#17838) --- Mathlib/Topology/ContinuousMap/Algebra.lean | 12 +++++++ Mathlib/Topology/ContinuousMap/Bounded.lean | 35 +++++++++++++++++++++ 2 files changed, 47 insertions(+) diff --git a/Mathlib/Topology/ContinuousMap/Algebra.lean b/Mathlib/Topology/ContinuousMap/Algebra.lean index 59525a680230d..f183492f6c269 100644 --- a/Mathlib/Topology/ContinuousMap/Algebra.lean +++ b/Mathlib/Topology/ContinuousMap/Algebra.lean @@ -581,6 +581,18 @@ instance [SMul R M] [ContinuousConstSMul R M] [SMul R₁ M] [ContinuousConstSMul instance [SMul R M] [SMul Rᵐᵒᵖ M] [ContinuousConstSMul R M] [IsCentralScalar R M] : IsCentralScalar R C(α, M) where op_smul_eq_smul _ _ := ext fun _ => op_smul_eq_smul _ _ +instance [SMul R M] [ContinuousConstSMul R M] [Mul M] [ContinuousMul M] [IsScalarTower R M M] : + IsScalarTower R C(α, M) C(α, M) where + smul_assoc _ _ _ := ext fun _ => smul_mul_assoc .. + +instance [SMul R M] [ContinuousConstSMul R M] [Mul M] [ContinuousMul M] [SMulCommClass R M M] : + SMulCommClass R C(α, M) C(α, M) where + smul_comm _ _ _ := ext fun _ => (mul_smul_comm ..).symm + +instance [SMul R M] [ContinuousConstSMul R M] [Mul M] [ContinuousMul M] [SMulCommClass M R M] : + SMulCommClass C(α, M) R C(α, M) where + smul_comm _ _ _ := ext fun _ => smul_comm (_ : M) .. + instance [Monoid R] [MulAction R M] [ContinuousConstSMul R M] : MulAction R C(α, M) := Function.Injective.mulAction _ coe_injective coe_smul diff --git a/Mathlib/Topology/ContinuousMap/Bounded.lean b/Mathlib/Topology/ContinuousMap/Bounded.lean index a6d612f35ff68..c85a6b15c3a3a 100644 --- a/Mathlib/Topology/ContinuousMap/Bounded.lean +++ b/Mathlib/Topology/ContinuousMap/Bounded.lean @@ -1013,6 +1013,16 @@ theorem coe_smul (c : 𝕜) (f : α →ᵇ β) : ⇑(c • f) = fun x => c • f theorem smul_apply (c : 𝕜) (f : α →ᵇ β) (x : α) : (c • f) x = c • f x := rfl +instance instIsScalarTower {𝕜' : Type*} [PseudoMetricSpace 𝕜'] [Zero 𝕜'] [SMul 𝕜' β] + [BoundedSMul 𝕜' β] [SMul 𝕜' 𝕜] [IsScalarTower 𝕜' 𝕜 β] : + IsScalarTower 𝕜' 𝕜 (α →ᵇ β) where + smul_assoc _ _ _ := ext fun _ ↦ smul_assoc .. + +instance instSMulCommClass {𝕜' : Type*} [PseudoMetricSpace 𝕜'] [Zero 𝕜'] [SMul 𝕜' β] + [BoundedSMul 𝕜' β] [SMulCommClass 𝕜' 𝕜 β] : + SMulCommClass 𝕜' 𝕜 (α →ᵇ β) where + smul_comm _ _ _ := ext fun _ ↦ smul_comm .. + instance instIsCentralScalar [SMul 𝕜ᵐᵒᵖ β] [IsCentralScalar 𝕜 β] : IsCentralScalar 𝕜 (α →ᵇ β) where op_smul_eq_smul _ _ := ext fun _ => op_smul_eq_smul _ _ @@ -1162,10 +1172,18 @@ instance instNonUnitalSeminormedRing : NonUnitalSeminormedRing (α →ᵇ R) := end Seminormed +instance instNonUnitalSeminormedCommRing [NonUnitalSeminormedCommRing R] : + NonUnitalSeminormedCommRing (α →ᵇ R) where + mul_comm _ _ := ext fun _ ↦ mul_comm .. + instance instNonUnitalNormedRing [NonUnitalNormedRing R] : NonUnitalNormedRing (α →ᵇ R) where __ := instNonUnitalSeminormedRing __ := instNormedAddCommGroup +instance instNonUnitalNormedCommRing [NonUnitalNormedCommRing R] : + NonUnitalNormedCommRing (α →ᵇ R) where + mul_comm := mul_comm + end NonUnital section Seminormed @@ -1246,6 +1264,23 @@ instance instNormedCommRing [NormedCommRing R] : NormedCommRing (α →ᵇ R) wh end NormedCommRing +section NonUnitalAlgebra + +-- these hypotheses could be generalized if we generalize `BoundedSMul` to `Bornology`. +variable {𝕜 : Type*} [PseudoMetricSpace 𝕜] [TopologicalSpace α] [NonUnitalSeminormedRing β] +variable [Zero 𝕜] [SMul 𝕜 β] [BoundedSMul 𝕜 β] + +instance [IsScalarTower 𝕜 β β] : IsScalarTower 𝕜 (α →ᵇ β) (α →ᵇ β) where + smul_assoc _ _ _ := ext fun _ ↦ smul_mul_assoc .. + +instance [SMulCommClass 𝕜 β β] : SMulCommClass 𝕜 (α →ᵇ β) (α →ᵇ β) where + smul_comm _ _ _ := ext fun _ ↦ (mul_smul_comm ..).symm + +instance [SMulCommClass 𝕜 β β] : SMulCommClass (α →ᵇ β) 𝕜 (α →ᵇ β) where + smul_comm _ _ _ := ext fun _ ↦ mul_smul_comm .. + +end NonUnitalAlgebra + section NormedAlgebra /-! From 4ed749a9bb0fca14b6bb4da5b5bf0a0825c4f08b Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Mon, 21 Oct 2024 15:16:22 +0000 Subject: [PATCH 242/505] RFC chore(GroupTheory/MonoidLocalization): un-`@[simp]` `Localization.mk_eq_monoidOf_mk'` (#18015) This lemma feels somewhat close to exposing an implementation detail to me, although it can be argued that the right hand side is indeed a bit simpler since it definitely uses more simple machinery. Again the balance is somewhat even between removing now-unnecessary `simp [-Localization.mk_eq_monoidOf_mk']` and adding new `simp [Localization.mk_eq_monoidOf_mk']`. Zulip thread: https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/RFC.3A.20un-.60simp.60ing.20.60normalize_apply.60 --- Mathlib/FieldTheory/RatFunc/Basic.lean | 2 +- Mathlib/GroupTheory/MonoidLocalization/Away.lean | 4 ++-- Mathlib/GroupTheory/MonoidLocalization/Basic.lean | 4 ++-- Mathlib/GroupTheory/MonoidLocalization/Order.lean | 2 +- Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean | 2 +- 5 files changed, 7 insertions(+), 7 deletions(-) diff --git a/Mathlib/FieldTheory/RatFunc/Basic.lean b/Mathlib/FieldTheory/RatFunc/Basic.lean index 4e0129beb6355..0aedae6b5183d 100644 --- a/Mathlib/FieldTheory/RatFunc/Basic.lean +++ b/Mathlib/FieldTheory/RatFunc/Basic.lean @@ -364,7 +364,7 @@ def map [MonoidHomClass F R[X] S[X]] (φ : F) (hφ : R[X]⁰ ≤ S[X]⁰.comap map_one' := by beta_reduce -- Porting note(#12129): force the function to be applied rw [← ofFractionRing_one, ← Localization.mk_one, liftOn_ofFractionRing_mk, dif_pos] - · simpa using ofFractionRing_one + · simpa [Localization.mk_eq_monoidOf_mk'] using ofFractionRing_one · simpa using Submonoid.one_mem _ map_mul' x y := by beta_reduce -- Porting note(#12129): force the function to be applied diff --git a/Mathlib/GroupTheory/MonoidLocalization/Away.lean b/Mathlib/GroupTheory/MonoidLocalization/Away.lean index 7bf763ad98465..e64618255bc9d 100644 --- a/Mathlib/GroupTheory/MonoidLocalization/Away.lean +++ b/Mathlib/GroupTheory/MonoidLocalization/Away.lean @@ -152,9 +152,9 @@ class of `(y, 0)` in the Localization of `M` at the Submonoid generated by `x`." abbrev Away.monoidOf : Submonoid.LocalizationMap.AwayMap x (Away x) := Localization.monoidOf (Submonoid.powers x) --- @[simp] -- Porting note (#10618): simp can prove thisrove this @[to_additive] -theorem Away.mk_eq_monoidOf_mk' : mk = (Away.monoidOf x).mk' := by simp +theorem Away.mk_eq_monoidOf_mk' : mk = (Away.monoidOf x).mk' := by + simp [Localization.mk_eq_monoidOf_mk'] /-- Given `x : M` and a Localization map `f : M →* N` away from `x`, we get an isomorphism between the Localization of `M` at the Submonoid generated by `x` as a quotient type and `N`. -/ diff --git a/Mathlib/GroupTheory/MonoidLocalization/Basic.lean b/Mathlib/GroupTheory/MonoidLocalization/Basic.lean index a05b795f2c3c4..5c49b43e602bb 100644 --- a/Mathlib/GroupTheory/MonoidLocalization/Basic.lean +++ b/Mathlib/GroupTheory/MonoidLocalization/Basic.lean @@ -1230,7 +1230,7 @@ theorem mk_eq_monoidOf_mk'_apply (x y) : mk x y = (monoidOf S).mk' x y := conv => rhs; rw [← mul_one 1]; rw [← mul_one x] exact mk_eq_mk_iff.2 (Con.symm _ <| (Localization.r S).mul (Con.refl _ (x, 1)) <| one_rel _) -@[to_additive (attr := simp)] +@[to_additive] theorem mk_eq_monoidOf_mk' : mk = (monoidOf S).mk' := funext fun _ ↦ funext fun _ ↦ mk_eq_monoidOf_mk'_apply _ _ @@ -1298,7 +1298,7 @@ variable {α : Type*} [CancelCommMonoid α] {s : Submonoid α} {a₁ b₁ : α} @[to_additive] theorem mk_left_injective (b : s) : Injective fun a => mk a b := fun c d h => by - simpa [-mk_eq_monoidOf_mk', mk_eq_mk_iff, r_iff_exists] using h + simpa [mk_eq_mk_iff, r_iff_exists] using h @[to_additive] theorem mk_eq_mk_iff' : mk a₁ a₂ = mk b₁ b₂ ↔ ↑b₂ * a₁ = a₂ * b₁ := by diff --git a/Mathlib/GroupTheory/MonoidLocalization/Order.lean b/Mathlib/GroupTheory/MonoidLocalization/Order.lean index 9dbdafbb9a75a..50502f62e0f37 100644 --- a/Mathlib/GroupTheory/MonoidLocalization/Order.lean +++ b/Mathlib/GroupTheory/MonoidLocalization/Order.lean @@ -106,7 +106,7 @@ sending `a` to `a - b`."] def mkOrderEmbedding (b : s) : α ↪o Localization s where toFun a := mk a b inj' := mk_left_injective _ - map_rel_iff' {a b} := by simp [-mk_eq_monoidOf_mk', mk_le_mk] + map_rel_iff' {a b} := by simp [mk_le_mk] end OrderedCancelCommMonoid diff --git a/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean b/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean index 762d89039a33b..f7b816f728486 100644 --- a/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean +++ b/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean @@ -528,7 +528,7 @@ theorem isUnit_iff_isUnit_val (f : HomogeneousLocalization.AtPrime 𝒜 𝔭) : (hc ▸ Ideal.mul_mem_left _ c.1 (Ideal.mul_mem_right b _ h)) refine isUnit_of_mul_eq_one _ (Quotient.mk'' ⟨f.1, f.3, f.2, this⟩) ?_ rw [← mk_mul, ext_iff_val, val_mk] - simp [mul_comm f.den.1] + simp [mul_comm f.den.1, Localization.mk_eq_monoidOf_mk'] instance : Nontrivial (HomogeneousLocalization.AtPrime 𝒜 𝔭) := ⟨⟨0, 1, fun r => by simp [ext_iff_val, val_zero, val_one, zero_ne_one] at r⟩⟩ From 0cb54c9f7ec15f9e156f7a3d9ffe4ace90e41f86 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Mon, 21 Oct 2024 15:44:23 +0000 Subject: [PATCH 243/505] feat(AlgebraicGeometry): dominant morphisms of schemes (#17961) Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- .../Morphisms/Constructors.lean | 20 +++++- .../Morphisms/UnderlyingMap.lean | 69 ++++++++++++++----- Mathlib/Topology/LocalAtTarget.lean | 12 ++++ 3 files changed, 81 insertions(+), 20 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Constructors.lean b/Mathlib/AlgebraicGeometry/Morphisms/Constructors.lean index 6cfb3141006be..bb4b68e99832c 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Constructors.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Constructors.lean @@ -222,8 +222,8 @@ lemma topologically_respectsIso we may check the corresponding properties on topological spaces. -/ lemma topologically_isLocalAtTarget [(topologically P).RespectsIso] - (hP₂ : ∀ {α β : Type u} [TopologicalSpace α] [TopologicalSpace β] (f : α → β) (s : Set β), - P f → P (s.restrictPreimage f)) + (hP₂ : ∀ {α β : Type u} [TopologicalSpace α] [TopologicalSpace β] (f : α → β) (s : Set β) + (_ : Continuous f) (_ : IsOpen s), P f → P (s.restrictPreimage f)) (hP₃ : ∀ {α β : Type u} [TopologicalSpace α] [TopologicalSpace β] (f : α → β) {ι : Type u} (U : ι → TopologicalSpace.Opens β) (_ : iSup U = ⊤) (_ : Continuous f), (∀ i, P ((U i).carrier.restrictPreimage f)) → P f) : @@ -231,12 +231,26 @@ lemma topologically_isLocalAtTarget apply IsLocalAtTarget.mk' · intro X Y f U hf simp_rw [topologically, morphismRestrict_base] - exact hP₂ f.base U.carrier hf + exact hP₂ f.base U.carrier f.base.2 U.2 hf · intro X Y f ι U hU hf apply hP₃ f.base U hU f.base.continuous fun i ↦ ?_ rw [← morphismRestrict_base] exact hf i +/-- A variant of `topologically_isLocalAtTarget` +that takes one iff statement instead of two implications. -/ +lemma topologically_isLocalAtTarget' + [(topologically P).RespectsIso] + (hP : ∀ {α β : Type u} [TopologicalSpace α] [TopologicalSpace β] (f : α → β) {ι : Type u} + (U : ι → TopologicalSpace.Opens β) (_ : iSup U = ⊤) (_ : Continuous f), + P f ↔ (∀ i, P ((U i).carrier.restrictPreimage f))) : + IsLocalAtTarget (topologically P) := by + refine topologically_isLocalAtTarget P ?_ (fun f _ U hU hU' ↦ (hP f U hU hU').mpr) + introv hf hs H + refine by simpa using (hP f (![⊤, Opens.mk s hs] ∘ Equiv.ulift) ?_ hf).mp H ⟨1⟩ + rw [← top_le_iff] + exact le_iSup (![⊤, Opens.mk s hs] ∘ Equiv.ulift) ⟨0⟩ + end Topologically /-- `stalkwise P` holds for a morphism if all stalks satisfy `P`. -/ diff --git a/Mathlib/AlgebraicGeometry/Morphisms/UnderlyingMap.lean b/Mathlib/AlgebraicGeometry/Morphisms/UnderlyingMap.lean index ed2f3268b2efc..3d31004a41fb1 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/UnderlyingMap.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/UnderlyingMap.lean @@ -20,6 +20,7 @@ of the underlying map of topological spaces, including - `Embedding` - `IsOpenEmbedding` - `IsClosedEmbedding` +- `DenseRange` (`IsDominant`) -/ @@ -37,7 +38,7 @@ instance : MorphismProperty.RespectsIso (topologically Function.Injective) := topologically_respectsIso _ (fun e ↦ e.injective) (fun _ _ hf hg ↦ hg.comp hf) instance injective_isLocalAtTarget : IsLocalAtTarget (topologically Function.Injective) := by - refine topologically_isLocalAtTarget _ (fun _ s h ↦ h.restrictPreimage s) + refine topologically_isLocalAtTarget _ (fun _ s _ _ h ↦ h.restrictPreimage s) fun f ι U H _ hf x₁ x₂ e ↦ ?_ obtain ⟨i, hxi⟩ : ∃ i, f x₁ ∈ U i := by simpa using congr(f x₁ ∈ $H) exact congr(($(@hf i ⟨x₁, hxi⟩ ⟨x₂, show f x₂ ∈ U i from e ▸ hxi⟩ (Subtype.ext e))).1) @@ -76,7 +77,8 @@ instance : MorphismProperty.RespectsIso @Surjective := instance surjective_isLocalAtTarget : IsLocalAtTarget @Surjective := by have : MorphismProperty.RespectsIso @Surjective := inferInstance rw [surjective_eq_topologically] at this ⊢ - refine topologically_isLocalAtTarget _ (fun _ s h ↦ h.restrictPreimage s) fun f ι U H _ hf x ↦ ?_ + refine topologically_isLocalAtTarget _ (fun _ s _ _ h ↦ h.restrictPreimage s) ?_ + intro α β _ _ f ι U H _ hf x obtain ⟨i, hxi⟩ : ∃ i, x ∈ U i := by simpa using congr(x ∈ $H) obtain ⟨⟨y, _⟩, hy⟩ := hf i ⟨x, hxi⟩ exact ⟨y, congr(($hy).1)⟩ @@ -89,9 +91,7 @@ instance : (topologically IsOpenMap).RespectsIso := topologically_respectsIso _ (fun e ↦ e.isOpenMap) (fun _ _ hf hg ↦ hg.comp hf) instance isOpenMap_isLocalAtTarget : IsLocalAtTarget (topologically IsOpenMap) := - topologically_isLocalAtTarget _ - (fun _ s hf ↦ hf.restrictPreimage s) - (fun _ _ _ hU _ hf ↦ (isOpenMap_iff_isOpenMap_of_iSup_eq_top hU).mpr hf) + topologically_isLocalAtTarget' _ fun _ _ _ hU _ ↦ isOpenMap_iff_isOpenMap_of_iSup_eq_top hU end IsOpenMap @@ -101,9 +101,7 @@ instance : (topologically IsClosedMap).RespectsIso := topologically_respectsIso _ (fun e ↦ e.isClosedMap) (fun _ _ hf hg ↦ hg.comp hf) instance isClosedMap_isLocalAtTarget : IsLocalAtTarget (topologically IsClosedMap) := - topologically_isLocalAtTarget _ - (fun _ s hf ↦ hf.restrictPreimage s) - (fun _ _ _ hU _ hf ↦ (isClosedMap_iff_isClosedMap_of_iSup_eq_top hU).mpr hf) + topologically_isLocalAtTarget' _ fun _ _ _ hU _ ↦ isClosedMap_iff_isClosedMap_of_iSup_eq_top hU end IsClosedMap @@ -113,9 +111,7 @@ instance : (topologically Embedding).RespectsIso := topologically_respectsIso _ (fun e ↦ e.embedding) (fun _ _ hf hg ↦ hg.comp hf) instance embedding_isLocalAtTarget : IsLocalAtTarget (topologically Embedding) := - topologically_isLocalAtTarget _ - (fun _ s hf ↦ hf.restrictPreimage s) - (fun _ _ _ hU hfcont hf ↦ (embedding_iff_embedding_of_iSup_eq_top hU hfcont).mpr hf) + topologically_isLocalAtTarget' _ fun _ _ _ ↦ embedding_iff_embedding_of_iSup_eq_top end Embedding @@ -125,9 +121,7 @@ instance : (topologically IsOpenEmbedding).RespectsIso := topologically_respectsIso _ (fun e ↦ e.isOpenEmbedding) (fun _ _ hf hg ↦ hg.comp hf) instance isOpenEmbedding_isLocalAtTarget : IsLocalAtTarget (topologically IsOpenEmbedding) := - topologically_isLocalAtTarget _ - (fun _ s hf ↦ hf.restrictPreimage s) - (fun _ _ _ hU hfcont hf ↦ (isOpenEmbedding_iff_isOpenEmbedding_of_iSup_eq_top hU hfcont).mpr hf) + topologically_isLocalAtTarget' _ fun _ _ _ ↦ isOpenEmbedding_iff_isOpenEmbedding_of_iSup_eq_top end IsOpenEmbedding @@ -137,10 +131,51 @@ instance : (topologically IsClosedEmbedding).RespectsIso := topologically_respectsIso _ (fun e ↦ e.isClosedEmbedding) (fun _ _ hf hg ↦ hg.comp hf) instance isClosedEmbedding_isLocalAtTarget : IsLocalAtTarget (topologically IsClosedEmbedding) := - topologically_isLocalAtTarget _ - (fun _ s hf ↦ hf.restrictPreimage s) - (fun _ _ _ hU hfcont ↦ (isClosedEmbedding_iff_isClosedEmbedding_of_iSup_eq_top hU hfcont).mpr) + topologically_isLocalAtTarget' _ + fun _ _ _ ↦ isClosedEmbedding_iff_isClosedEmbedding_of_iSup_eq_top end IsClosedEmbedding +section IsDominant + +variable {X Y Z : Scheme.{u}} (f : X ⟶ Y) (g : Y ⟶ Z) + +/-- A morphism of schemes is dominant if the underlying map has dense range. -/ +@[mk_iff] +class IsDominant : Prop where + denseRange : DenseRange f.base + +lemma dominant_eq_topologically : + @IsDominant = topologically DenseRange := by ext; exact isDominant_iff _ + +lemma Scheme.Hom.denseRange (f : X.Hom Y) [IsDominant f] : DenseRange f.base := + IsDominant.denseRange + +instance (priority := 100) [Surjective f] : IsDominant f := ⟨f.surjective.denseRange⟩ + +instance [IsDominant f] [IsDominant g] : IsDominant (f ≫ g) := + ⟨g.denseRange.comp f.denseRange g.base.2⟩ + +instance : MorphismProperty.IsMultiplicative @IsDominant where + id_mem := fun _ ↦ inferInstance + comp_mem := fun _ _ _ _ ↦ inferInstance + +lemma IsDominant.of_comp [H : IsDominant (f ≫ g)] : IsDominant g := by + rw [isDominant_iff, denseRange_iff_closure_range, ← Set.univ_subset_iff] at H ⊢ + exact H.trans (closure_mono (Set.range_comp_subset_range f.base g.base)) + +lemma IsDominant.comp_iff [IsDominant f] : IsDominant (f ≫ g) ↔ IsDominant g := + ⟨fun _ ↦ of_comp f g, fun _ ↦ inferInstance⟩ + +instance IsDominant.respectsIso : MorphismProperty.RespectsIso @IsDominant := + MorphismProperty.respectsIso_of_isStableUnderComposition fun _ _ f (_ : IsIso f) ↦ inferInstance + +instance IsDominant.isLocalAtTarget : IsLocalAtTarget @IsDominant := + have : MorphismProperty.RespectsIso (topologically DenseRange) := + dominant_eq_topologically ▸ IsDominant.respectsIso + dominant_eq_topologically ▸ topologically_isLocalAtTarget' DenseRange + fun _ _ _ hU _ ↦ denseRange_iff_denseRange_of_iSup_eq_top hU + +end IsDominant + end AlgebraicGeometry diff --git a/Mathlib/Topology/LocalAtTarget.lean b/Mathlib/Topology/LocalAtTarget.lean index c1e9f1b8153f2..73a4072157b9c 100644 --- a/Mathlib/Topology/LocalAtTarget.lean +++ b/Mathlib/Topology/LocalAtTarget.lean @@ -199,6 +199,18 @@ theorem isClosedEmbedding_iff_isClosedEmbedding_of_iSup_eq_top (h : Continuous f · simp_rw [Set.range_restrictPreimage] apply isClosed_iff_coe_preimage_of_iSup_eq_top hU +omit [TopologicalSpace α] in +theorem denseRange_iff_denseRange_of_iSup_eq_top : + DenseRange f ↔ ∀ i, DenseRange ((U i).1.restrictPreimage f) := by + simp_rw [denseRange_iff_closure_range, Set.range_restrictPreimage, + ← (U _).2.isOpenEmbedding_subtypeVal.isOpenMap.preimage_closure_eq_closure_preimage + continuous_subtype_val] + simp only [Opens.carrier_eq_coe, SetLike.coe_sort_coe, preimage_eq_univ_iff, + Subtype.range_coe_subtype, SetLike.mem_coe] + rw [← iUnion_subset_iff, ← Set.univ_subset_iff, iff_iff_eq] + congr 1 + simpa using congr(($hU).1).symm + @[deprecated (since := "2024-10-20")] alias closedEmbedding_iff_closedEmbedding_of_iSup_eq_top := isClosedEmbedding_iff_isClosedEmbedding_of_iSup_eq_top From e3e46c381d49179c3818135c0a888bc69a958de3 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 21 Oct 2024 15:51:40 +0000 Subject: [PATCH 244/505] =?UTF-8?q?feat:=20`=E2=80=96a=20*=20b=20*=20c?= =?UTF-8?q?=E2=80=96=20=E2=89=A4=20=E2=80=96a=E2=80=96=20*=20=E2=80=96b?= =?UTF-8?q?=E2=80=96=20*=20=E2=80=96c=E2=80=96`=20(#18005)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit and prime the existing lemmas for a norm on a multiplicative group to make space for the ring ones. From GrowthInGroups --- Mathlib/Analysis/Analytic/Basic.lean | 2 +- Mathlib/Analysis/CStarAlgebra/Module/Defs.lean | 2 +- Mathlib/Analysis/Calculus/FDeriv/Measurable.lean | 4 ++-- Mathlib/Analysis/Calculus/Rademacher.lean | 2 +- Mathlib/Analysis/Convex/Uniform.lean | 4 ++-- Mathlib/Analysis/Normed/Field/Basic.lean | 13 ++++++++++++- Mathlib/Analysis/Normed/Group/Basic.lean | 15 ++++++++------- Mathlib/Analysis/Normed/Group/Pointwise.lean | 2 +- Mathlib/Probability/StrongLaw.lean | 2 +- 9 files changed, 29 insertions(+), 17 deletions(-) diff --git a/Mathlib/Analysis/Analytic/Basic.lean b/Mathlib/Analysis/Analytic/Basic.lean index 35af733c58d75..ecc9922cff892 100644 --- a/Mathlib/Analysis/Analytic/Basic.lean +++ b/Mathlib/Analysis/Analytic/Basic.lean @@ -890,7 +890,7 @@ theorem HasFPowerSeriesWithinOnBall.tendsto_partialSum_prod {y : E} simp only [FormalMultilinearSeries.partialSum] abel _ ≤ ‖p.partialSum k z - p.partialSum k y‖ + ‖∑ i ∈ Ico k n, p i (fun _ ↦ z)‖ - + ‖- ∑ i ∈ Ico k n, p i (fun _ ↦ y)‖ := norm_add₃_le _ _ _ + + ‖- ∑ i ∈ Ico k n, p i (fun _ ↦ y)‖ := norm_add₃_le _ ≤ ε / 4 + ε / 4 + ε / 4 := by gcongr · exact I _ h'z diff --git a/Mathlib/Analysis/CStarAlgebra/Module/Defs.lean b/Mathlib/Analysis/CStarAlgebra/Module/Defs.lean index c848155195473..2b601c936ad06 100644 --- a/Mathlib/Analysis/CStarAlgebra/Module/Defs.lean +++ b/Mathlib/Analysis/CStarAlgebra/Module/Defs.lean @@ -241,7 +241,7 @@ protected lemma norm_triangle [CompleteSpace A] (x y : E) : ‖x + y‖ ≤ ‖x calc _ ≤ ‖⟪x, x⟫ + ⟪y, x⟫‖ + ‖⟪x, y⟫‖ + ‖⟪y, y⟫‖ := by simp only [norm_eq_sqrt_norm_inner_self, inner_add_right, inner_add_left, ← add_assoc, norm_nonneg, Real.sq_sqrt] - exact norm_add₃_le _ _ _ + exact norm_add₃_le _ ≤ ‖⟪x, x⟫‖ + ‖⟪y, x⟫‖ + ‖⟪x, y⟫‖ + ‖⟪y, y⟫‖ := by gcongr; exact norm_add_le _ _ _ ≤ ‖⟪x, x⟫‖ + ‖y‖ * ‖x‖ + ‖x‖ * ‖y‖ + ‖⟪y, y⟫‖ := by gcongr <;> exact norm_inner_le E _ = ‖x‖ ^ 2 + ‖y‖ * ‖x‖ + ‖x‖ * ‖y‖ + ‖y‖ ^ 2 := by diff --git a/Mathlib/Analysis/Calculus/FDeriv/Measurable.lean b/Mathlib/Analysis/Calculus/FDeriv/Measurable.lean index 5801ab5775fa7..10ba2121749bd 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Measurable.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Measurable.lean @@ -259,7 +259,7 @@ theorem D_subset_differentiable_set {K : Set (E →L[𝕜] F)} (hK : IsComplete ‖L e p q - L e p r + (L e p r - L e' p' r) + (L e' p' r - L e' p' q')‖ := by congr 1; abel _ ≤ ‖L e p q - L e p r‖ + ‖L e p r - L e' p' r‖ + ‖L e' p' r - L e' p' q'‖ := - norm_add₃_le _ _ _ + norm_add₃_le _ ≤ 4 * ‖c‖ * (1 / 2) ^ e + 4 * ‖c‖ * (1 / 2) ^ e + 4 * ‖c‖ * (1 / 2) ^ e := by gcongr _ = 12 * ‖c‖ * (1 / 2) ^ e := by ring /- For definiteness, use `L0 e = L e (n e) (n e)`, to have a single sequence. We claim that this @@ -843,7 +843,7 @@ lemma isOpen_A_with_param {r s : ℝ} (hf : Continuous f.uncurry) (L : E →L[ calc ‖f a' z - f a' y - (L z - L y)‖ = ‖(f a' z - f a z) + (f a y - f a' y) + (f a z - f a y - (L z - L y))‖ := by congr; abel - _ ≤ ‖f a' z - f a z‖ + ‖f a y - f a' y‖ + ‖f a z - f a y - (L z - L y)‖ := norm_add₃_le _ _ _ + _ ≤ ‖f a' z - f a z‖ + ‖f a y - f a' y‖ + ‖f a z - f a y - (L z - L y)‖ := norm_add₃_le _ ≤ ε + ε + b := by gcongr · rw [← dist_eq_norm] diff --git a/Mathlib/Analysis/Calculus/Rademacher.lean b/Mathlib/Analysis/Calculus/Rademacher.lean index 67771abec0dce..28ad0559741a1 100644 --- a/Mathlib/Analysis/Calculus/Rademacher.lean +++ b/Mathlib/Analysis/Calculus/Rademacher.lean @@ -298,7 +298,7 @@ theorem hasFderivAt_of_hasLineDerivAt_of_closure _ = ‖(f (x + ρ • w) - f (x + ρ • y)) + (ρ • L y - ρ • L w) + (f (x + ρ • y) - f x - ρ • L y)‖ := by congr; abel _ ≤ ‖f (x + ρ • w) - f (x + ρ • y)‖ + ‖ρ • L y - ρ • L w‖ - + ‖f (x + ρ • y) - f x - ρ • L y‖ := norm_add₃_le _ _ _ + + ‖f (x + ρ • y) - f x - ρ • L y‖ := norm_add₃_le _ ≤ C * ‖(x + ρ • w) - (x + ρ • y)‖ + ρ * (‖L‖ * ‖y - w‖) + δ * ρ := by gcongr · exact hf.norm_sub_le _ _ diff --git a/Mathlib/Analysis/Convex/Uniform.lean b/Mathlib/Analysis/Convex/Uniform.lean index 9e22c41ca4d16..152127c0a4736 100644 --- a/Mathlib/Analysis/Convex/Uniform.lean +++ b/Mathlib/Analysis/Convex/Uniform.lean @@ -87,12 +87,12 @@ theorem exists_forall_closed_ball_dist_add_le_two_sub (hε : 0 < ε) : _ ≤ _ := by have : ∀ x' y', x - y = x' - y' + (x - x') + (y' - y) := fun _ _ => by abel rw [sub_le_iff_le_add, norm_sub_rev _ x, ← add_assoc, this] - exact norm_add₃_le _ _ _ + exact norm_add₃_le calc ‖x + y‖ ≤ ‖x' + y'‖ + ‖x' - x‖ + ‖y' - y‖ := by have : ∀ x' y', x + y = x' + y' + (x - x') + (y - y') := fun _ _ => by abel rw [norm_sub_rev, norm_sub_rev y', this] - exact norm_add₃_le _ _ _ + exact norm_add₃_le _ ≤ 2 - δ + δ' + δ' := (add_le_add_three (h (h₁ _ hx') (h₁ _ hy') hxy') (h₂ _ hx hx'.le) (h₂ _ hy hy'.le)) _ ≤ 2 - δ' := by diff --git a/Mathlib/Analysis/Normed/Field/Basic.lean b/Mathlib/Analysis/Normed/Field/Basic.lean index 405b9ae12e4fa..346a70a6a9cc2 100644 --- a/Mathlib/Analysis/Normed/Field/Basic.lean +++ b/Mathlib/Analysis/Normed/Field/Basic.lean @@ -212,7 +212,7 @@ instance MulOpposite.normOneClass [SeminormedAddCommGroup α] [One α] [NormOneC section NonUnitalSeminormedRing -variable [NonUnitalSeminormedRing α] +variable [NonUnitalSeminormedRing α] {a a₁ a₂ b c : α} theorem norm_mul_le (a b : α) : ‖a * b‖ ≤ ‖a‖ * ‖b‖ := NonUnitalSeminormedRing.norm_mul _ _ @@ -221,6 +221,17 @@ theorem nnnorm_mul_le (a b : α) : ‖a * b‖₊ ≤ ‖a‖₊ * ‖b‖₊ := simpa only [← norm_toNNReal, ← Real.toNNReal_mul (norm_nonneg _)] using Real.toNNReal_mono (norm_mul_le _ _) +lemma norm_mul_le_of_le {r₁ r₂ : ℝ} (h₁ : ‖a₁‖ ≤ r₁) (h₂ : ‖a₂‖ ≤ r₂) : ‖a₁ * a₂‖ ≤ r₁ * r₂ := + (norm_mul_le ..).trans <| mul_le_mul h₁ h₂ (norm_nonneg _) ((norm_nonneg _).trans h₁) + +lemma nnnorm_mul_le_of_le {r₁ r₂ : ℝ≥0} (h₁ : ‖a₁‖₊ ≤ r₁) (h₂ : ‖a₂‖₊ ≤ r₂) : + ‖a₁ * a₂‖₊ ≤ r₁ * r₂ := (nnnorm_mul_le ..).trans <| mul_le_mul' h₁ h₂ + +lemma norm_mul₃_le : ‖a * b * c‖ ≤ ‖a‖ * ‖b‖ * ‖c‖ := norm_mul_le_of_le (norm_mul_le ..) le_rfl + +lemma nnnorm_mul₃_le : ‖a * b * c‖₊ ≤ ‖a‖₊ * ‖b‖₊ * ‖c‖₊ := + nnnorm_mul_le_of_le (norm_mul_le ..) le_rfl + theorem one_le_norm_one (β) [NormedRing β] [Nontrivial β] : 1 ≤ ‖(1 : β)‖ := (le_mul_iff_one_le_left <| norm_pos_iff.mpr (one_ne_zero : (1 : β) ≠ 0)).mp (by simpa only [mul_one] using norm_mul_le (1 : β) 1) diff --git a/Mathlib/Analysis/Normed/Group/Basic.lean b/Mathlib/Analysis/Normed/Group/Basic.lean index b3742ef73f777..9ff5a866463b2 100644 --- a/Mathlib/Analysis/Normed/Group/Basic.lean +++ b/Mathlib/Analysis/Normed/Group/Basic.lean @@ -338,7 +338,7 @@ abbrev GroupNorm.toNormedCommGroup [CommGroup E] (f : GroupNorm E) : NormedCommG section SeminormedGroup variable [SeminormedGroup E] [SeminormedGroup F] [SeminormedGroup G] {s : Set E} - {a a₁ a₂ b b₁ b₂ : E} {r r₁ r₂ : ℝ} + {a a₁ a₂ b b₁ b₂ c : E} {r r₁ r₂ : ℝ} @[to_additive] theorem dist_eq_norm_div (a b : E) : dist a b = ‖a / b‖ := @@ -387,13 +387,14 @@ theorem dist_mulIndicator (s t : Set α) (f : α → E) (x : α) : theorem norm_mul_le' (a b : E) : ‖a * b‖ ≤ ‖a‖ + ‖b‖ := by simpa [dist_eq_norm_div] using dist_triangle a 1 b⁻¹ -@[to_additive] -theorem norm_mul_le_of_le (h₁ : ‖a₁‖ ≤ r₁) (h₂ : ‖a₂‖ ≤ r₂) : ‖a₁ * a₂‖ ≤ r₁ + r₂ := +/-- **Triangle inequality** for the norm. -/ +@[to_additive norm_add_le_of_le "**Triangle inequality** for the norm."] +theorem norm_mul_le_of_le' (h₁ : ‖a₁‖ ≤ r₁) (h₂ : ‖a₂‖ ≤ r₂) : ‖a₁ * a₂‖ ≤ r₁ + r₂ := (norm_mul_le' a₁ a₂).trans <| add_le_add h₁ h₂ -@[to_additive norm_add₃_le] -theorem norm_mul₃_le (a b c : E) : ‖a * b * c‖ ≤ ‖a‖ + ‖b‖ + ‖c‖ := - norm_mul_le_of_le (norm_mul_le' _ _) le_rfl +/-- **Triangle inequality** for the norm. -/ +@[to_additive norm_add₃_le "**Triangle inequality** for the norm."] +lemma norm_mul₃_le' : ‖a * b * c‖ ≤ ‖a‖ + ‖b‖ + ‖c‖ := norm_mul_le_of_le' (norm_mul_le' _ _) le_rfl @[to_additive] lemma norm_div_le_norm_div_add_norm_div (a b c : E) : ‖a / c‖ ≤ ‖a / b‖ + ‖b / c‖ := by @@ -1049,7 +1050,7 @@ theorem norm_pow_le_mul_norm (n : ℕ) (a : E) : ‖a ^ n‖ ≤ n * ‖a‖ := induction n with | zero => simp | succ n ih => - simpa only [pow_succ, Nat.cast_succ, add_mul, one_mul] using norm_mul_le_of_le ih le_rfl + simpa only [pow_succ, Nat.cast_succ, add_mul, one_mul] using norm_mul_le_of_le' ih le_rfl @[to_additive nnnorm_nsmul_le] theorem nnnorm_pow_le_mul_norm (n : ℕ) (a : E) : ‖a ^ n‖₊ ≤ n * ‖a‖₊ := by diff --git a/Mathlib/Analysis/Normed/Group/Pointwise.lean b/Mathlib/Analysis/Normed/Group/Pointwise.lean index ae8862a8aa948..f06a7d00a7621 100644 --- a/Mathlib/Analysis/Normed/Group/Pointwise.lean +++ b/Mathlib/Analysis/Normed/Group/Pointwise.lean @@ -31,7 +31,7 @@ theorem Bornology.IsBounded.mul (hs : IsBounded s) (ht : IsBounded t) : IsBounde obtain ⟨Rt, hRt⟩ : ∃ R, ∀ x ∈ t, ‖x‖ ≤ R := ht.exists_norm_le' refine isBounded_iff_forall_norm_le'.2 ⟨Rs + Rt, ?_⟩ rintro z ⟨x, hx, y, hy, rfl⟩ - exact norm_mul_le_of_le (hRs x hx) (hRt y hy) + exact norm_mul_le_of_le' (hRs x hx) (hRt y hy) @[to_additive] theorem Bornology.IsBounded.of_mul (hst : IsBounded (s * t)) : IsBounded s ∨ IsBounded t := diff --git a/Mathlib/Probability/StrongLaw.lean b/Mathlib/Probability/StrongLaw.lean index 70fe8e747018d..fae0c8332ee2f 100644 --- a/Mathlib/Probability/StrongLaw.lean +++ b/Mathlib/Probability/StrongLaw.lean @@ -775,7 +775,7 @@ lemma strong_law_ae_of_measurable abel _ ≤ ‖(n : ℝ)⁻¹ • ∑ i ∈ Finset.range n, (X i ω - Y k i ω)‖ + ‖(n : ℝ)⁻¹ • ∑ i ∈ Finset.range n, Y k i ω - μ [Y k 0]‖ + ‖μ [Y k 0] - μ [X 0]‖ := - norm_add₃_le _ _ _ + norm_add₃_le _ ≤ (∑ i ∈ Finset.range n, ‖X i ω - Y k i ω‖) / n + δ + δ := by gcongr simp only [Function.comp_apply, norm_smul, norm_inv, RCLike.norm_natCast, From 340899ebd0cae121230bf6e70a6d1e7bcf330ed0 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 21 Oct 2024 16:18:10 +0000 Subject: [PATCH 245/505] chore(Pointwise/Finset): move `MonoidWithZero` and `Ring` results out (#17999) Move them to two new files `Algebra.GroupWithZero.Pointwise.Finset` and `Algebra.Ring.Pointwise.Finset` --- Mathlib.lean | 2 + .../Algebra/Group/Pointwise/Finset/Basic.lean | 260 +----------------- .../GroupWithZero/Pointwise/Finset.lean | 209 ++++++++++++++ Mathlib/Algebra/Ring/Pointwise/Finset.lean | 69 +++++ 4 files changed, 284 insertions(+), 256 deletions(-) create mode 100644 Mathlib/Algebra/GroupWithZero/Pointwise/Finset.lean create mode 100644 Mathlib/Algebra/Ring/Pointwise/Finset.lean diff --git a/Mathlib.lean b/Mathlib.lean index 1f124da7a5e7a..539841cfd97d9 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -329,6 +329,7 @@ import Mathlib.Algebra.GroupWithZero.NeZero import Mathlib.Algebra.GroupWithZero.NonZeroDivisors import Mathlib.Algebra.GroupWithZero.Opposite import Mathlib.Algebra.GroupWithZero.Pi +import Mathlib.Algebra.GroupWithZero.Pointwise.Finset import Mathlib.Algebra.GroupWithZero.Pointwise.Set.Basic import Mathlib.Algebra.GroupWithZero.Pointwise.Set.Card import Mathlib.Algebra.GroupWithZero.Prod @@ -796,6 +797,7 @@ import Mathlib.Algebra.Ring.NegOnePow import Mathlib.Algebra.Ring.Opposite import Mathlib.Algebra.Ring.Parity import Mathlib.Algebra.Ring.Pi +import Mathlib.Algebra.Ring.Pointwise.Finset import Mathlib.Algebra.Ring.Pointwise.Set import Mathlib.Algebra.Ring.Prod import Mathlib.Algebra.Ring.Rat diff --git a/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean b/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean index c51cd98506df2..22ddb26282b4f 100644 --- a/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean +++ b/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean @@ -5,8 +5,6 @@ Authors: Floris van Doorn, Yaël Dillies -/ import Mathlib.Algebra.BigOperators.Group.Finset import Mathlib.Algebra.Group.Action.Pi -import Mathlib.Algebra.Order.Ring.Nat -import Mathlib.Algebra.Ring.Pointwise.Set import Mathlib.Data.Finset.Density import Mathlib.Data.Finset.NAry import Mathlib.Data.Set.Pointwise.Finite @@ -55,7 +53,8 @@ finset multiplication, finset addition, pointwise addition, pointwise multiplica pointwise subtraction -/ - +-- TODO +-- assert_not_exists MonoidWithZero assert_not_exists Cardinal open Function MulOpposite @@ -915,58 +914,7 @@ end DivisionMonoid protected def divisionCommMonoid [DivisionCommMonoid α] : DivisionCommMonoid (Finset α) := coe_injective.divisionCommMonoid _ coe_one coe_mul coe_inv coe_div coe_pow coe_zpow -/-- `Finset α` has distributive negation if `α` has. -/ -protected def distribNeg [Mul α] [HasDistribNeg α] : HasDistribNeg (Finset α) := - coe_injective.hasDistribNeg _ coe_neg coe_mul - -scoped[Pointwise] - attribute [instance] Finset.divisionCommMonoid Finset.subtractionCommMonoid Finset.distribNeg - -section Distrib - -variable [Distrib α] (s t u : Finset α) - -/-! -Note that `Finset α` is not a `Distrib` because `s * t + s * u` has cross terms that `s * (t + u)` -lacks. - -```lean --- {10, 16, 18, 20, 8, 9} -#eval {1, 2} * ({3, 4} + {5, 6} : Finset ℕ) - --- {10, 11, 12, 13, 14, 15, 16, 18, 20, 8, 9} -#eval ({1, 2} : Finset ℕ) * {3, 4} + {1, 2} * {5, 6} -``` --/ - - -theorem mul_add_subset : s * (t + u) ⊆ s * t + s * u := - image₂_distrib_subset_left mul_add - -theorem add_mul_subset : (s + t) * u ⊆ s * u + t * u := - image₂_distrib_subset_right add_mul - -end Distrib - -section MulZeroClass - -variable [MulZeroClass α] {s : Finset α} - -/-! Note that `Finset` is not a `MulZeroClass` because `0 * ∅ ≠ 0`. -/ - - -theorem mul_zero_subset (s : Finset α) : s * 0 ⊆ 0 := by simp [subset_iff, mem_mul] - -theorem zero_mul_subset (s : Finset α) : 0 * s ⊆ 0 := by simp [subset_iff, mem_mul] - -theorem Nonempty.mul_zero (hs : s.Nonempty) : s * 0 = 0 := - s.mul_zero_subset.antisymm <| by simpa [mem_mul] using hs - -theorem Nonempty.zero_mul (hs : s.Nonempty) : 0 * s = 0 := - s.zero_mul_subset.antisymm <| by simpa [mem_mul] using hs - -end MulZeroClass - +scoped[Pointwise] attribute [instance] Finset.divisionCommMonoid Finset.subtractionCommMonoid section Group variable [Group α] [DivisionMonoid β] [FunLike F α β] [MonoidHomClass F α β] @@ -1029,22 +977,6 @@ theorem image_div : (s / t).image (f : α → β) = s.image f / t.image f := end Group -section GroupWithZero - -variable [GroupWithZero α] {s : Finset α} - -theorem div_zero_subset (s : Finset α) : s / 0 ⊆ 0 := by simp [subset_iff, mem_div] - -theorem zero_div_subset (s : Finset α) : 0 / s ⊆ 0 := by simp [subset_iff, mem_div] - -theorem Nonempty.div_zero (hs : s.Nonempty) : s / 0 = 0 := - s.div_zero_subset.antisymm <| by simpa [mem_div] using hs - -theorem Nonempty.zero_div (hs : s.Nonempty) : 0 / s = 0 := - s.zero_div_subset.antisymm <| by simpa [mem_div] using hs - -end GroupWithZero - end Instances section Group @@ -1475,48 +1407,6 @@ scoped[Pointwise] attribute [instance] Finset.mulActionFinset Finset.addActionFinset Finset.mulAction Finset.addAction -/-- If scalar multiplication by elements of `α` sends `(0 : β)` to zero, -then the same is true for `(0 : Finset β)`. -/ -protected def smulZeroClassFinset [Zero β] [SMulZeroClass α β] : - SMulZeroClass α (Finset β) := - coe_injective.smulZeroClass ⟨(↑), coe_zero⟩ coe_smul_finset - -scoped[Pointwise] attribute [instance] Finset.smulZeroClassFinset - -/-- If the scalar multiplication `(· • ·) : α → β → β` is distributive, -then so is `(· • ·) : α → Finset β → Finset β`. -/ -protected def distribSMulFinset [AddZeroClass β] [DistribSMul α β] : - DistribSMul α (Finset β) := - coe_injective.distribSMul coeAddMonoidHom coe_smul_finset - -scoped[Pointwise] attribute [instance] Finset.distribSMulFinset - -/-- A distributive multiplicative action of a monoid on an additive monoid `β` gives a distributive -multiplicative action on `Finset β`. -/ -protected def distribMulActionFinset [Monoid α] [AddMonoid β] [DistribMulAction α β] : - DistribMulAction α (Finset β) := - Function.Injective.distribMulAction coeAddMonoidHom coe_injective coe_smul_finset - -/-- A multiplicative action of a monoid on a monoid `β` gives a multiplicative action on `Set β`. -/ -protected def mulDistribMulActionFinset [Monoid α] [Monoid β] [MulDistribMulAction α β] : - MulDistribMulAction α (Finset β) := - Function.Injective.mulDistribMulAction coeMonoidHom coe_injective coe_smul_finset - -scoped[Pointwise] - attribute [instance] Finset.distribMulActionFinset Finset.mulDistribMulActionFinset - -instance [DecidableEq α] [Zero α] [Mul α] [NoZeroDivisors α] : NoZeroDivisors (Finset α) := - Function.Injective.noZeroDivisors (↑) coe_injective coe_zero coe_mul - -instance noZeroSMulDivisors [Zero α] [Zero β] [SMul α β] [NoZeroSMulDivisors α β] : - NoZeroSMulDivisors (Finset α) (Finset β) where - eq_zero_or_eq_zero_of_smul_eq_zero {s t} := by - exact_mod_cast eq_zero_or_eq_zero_of_smul_eq_zero (c := s.toSet) (x := t.toSet) - -instance noZeroSMulDivisors_finset [Zero α] [Zero β] [SMul α β] [NoZeroSMulDivisors α β] : - NoZeroSMulDivisors α (Finset β) := - Function.Injective.noZeroSMulDivisors (↑) coe_injective coe_zero coe_smul_finset - end Instances section SMul @@ -1753,148 +1643,6 @@ lemma inv_op_smul_finset_distrib (a : α) (s : Finset α) : (op a • s)⁻¹ = end Group -section SMulZeroClass - -variable [Zero β] [SMulZeroClass α β] [DecidableEq β] {s : Finset α} {t : Finset β} {a : α} - -theorem smul_zero_subset (s : Finset α) : s • (0 : Finset β) ⊆ 0 := by simp [subset_iff, mem_smul] - -theorem Nonempty.smul_zero (hs : s.Nonempty) : s • (0 : Finset β) = 0 := - s.smul_zero_subset.antisymm <| by simpa [mem_smul] using hs - -theorem zero_mem_smul_finset (h : (0 : β) ∈ t) : (0 : β) ∈ a • t := - mem_smul_finset.2 ⟨0, h, smul_zero _⟩ - -variable [Zero α] [NoZeroSMulDivisors α β] - -theorem zero_mem_smul_finset_iff (ha : a ≠ 0) : (0 : β) ∈ a • t ↔ (0 : β) ∈ t := by - rw [← mem_coe, coe_smul_finset, Set.zero_mem_smul_set_iff ha, mem_coe] - -end SMulZeroClass - -section SMulWithZero - -variable [Zero α] [Zero β] [SMulWithZero α β] [DecidableEq β] {s : Finset α} {t : Finset β} - -/-! -Note that we have neither `SMulWithZero α (Finset β)` nor `SMulWithZero (Finset α) (Finset β)` -because `0 • ∅ ≠ 0`. --/ - -lemma zero_smul_subset (t : Finset β) : (0 : Finset α) • t ⊆ 0 := by simp [subset_iff, mem_smul] - -lemma Nonempty.zero_smul (ht : t.Nonempty) : (0 : Finset α) • t = 0 := - t.zero_smul_subset.antisymm <| by simpa [mem_smul] using ht - -/-- A nonempty set is scaled by zero to the singleton set containing zero. -/ -@[simp] lemma zero_smul_finset {s : Finset β} (h : s.Nonempty) : (0 : α) • s = (0 : Finset β) := - coe_injective <| by simpa using @Set.zero_smul_set α _ _ _ _ _ h - -lemma zero_smul_finset_subset (s : Finset β) : (0 : α) • s ⊆ 0 := - image_subset_iff.2 fun x _ ↦ mem_zero.2 <| zero_smul α x - -variable [NoZeroSMulDivisors α β] - -lemma zero_mem_smul_iff : - (0 : β) ∈ s • t ↔ (0 : α) ∈ s ∧ t.Nonempty ∨ (0 : β) ∈ t ∧ s.Nonempty := by - rw [← mem_coe, coe_smul, Set.zero_mem_smul_iff]; rfl - -end SMulWithZero - -section GroupWithZero - -variable [DecidableEq β] [GroupWithZero α] [MulAction α β] {s t : Finset β} {a : α} {b : β} - -@[simp] -theorem smul_mem_smul_finset_iff₀ (ha : a ≠ 0) : a • b ∈ a • s ↔ b ∈ s := - smul_mem_smul_finset_iff (Units.mk0 a ha) - -theorem inv_smul_mem_iff₀ (ha : a ≠ 0) : a⁻¹ • b ∈ s ↔ b ∈ a • s := - show _ ↔ _ ∈ Units.mk0 a ha • _ from inv_smul_mem_iff - -theorem mem_inv_smul_finset_iff₀ (ha : a ≠ 0) : b ∈ a⁻¹ • s ↔ a • b ∈ s := - show _ ∈ (Units.mk0 a ha)⁻¹ • _ ↔ _ from mem_inv_smul_finset_iff - -@[simp] -theorem smul_finset_subset_smul_finset_iff₀ (ha : a ≠ 0) : a • s ⊆ a • t ↔ s ⊆ t := - show Units.mk0 a ha • _ ⊆ _ ↔ _ from smul_finset_subset_smul_finset_iff - -theorem smul_finset_subset_iff₀ (ha : a ≠ 0) : a • s ⊆ t ↔ s ⊆ a⁻¹ • t := - show Units.mk0 a ha • _ ⊆ _ ↔ _ from smul_finset_subset_iff - -theorem subset_smul_finset_iff₀ (ha : a ≠ 0) : s ⊆ a • t ↔ a⁻¹ • s ⊆ t := - show _ ⊆ Units.mk0 a ha • _ ↔ _ from subset_smul_finset_iff - -theorem smul_finset_inter₀ (ha : a ≠ 0) : a • (s ∩ t) = a • s ∩ a • t := - image_inter _ _ <| MulAction.injective₀ ha - -theorem smul_finset_sdiff₀ (ha : a ≠ 0) : a • (s \ t) = a • s \ a • t := - image_sdiff _ _ <| MulAction.injective₀ ha - -open scoped symmDiff in -theorem smul_finset_symmDiff₀ (ha : a ≠ 0) : a • s ∆ t = (a • s) ∆ (a • t) := - image_symmDiff _ _ <| MulAction.injective₀ ha - -lemma smul_finset_univ₀ [Fintype β] (ha : a ≠ 0) : a • (univ : Finset β) = univ := - coe_injective <| by push_cast; exact Set.smul_set_univ₀ ha - -theorem smul_univ₀ [Fintype β] {s : Finset α} (hs : ¬s ⊆ 0) : s • (univ : Finset β) = univ := - coe_injective <| by - rw [← coe_subset] at hs - push_cast at hs ⊢ - exact Set.smul_univ₀ hs - -lemma smul_univ₀' [Fintype β] {s : Finset α} (hs : s.Nontrivial) : s • (univ : Finset β) = univ := - coe_injective <| by push_cast; exact Set.smul_univ₀' hs - -variable [DecidableEq α] - -@[simp] protected lemma inv_zero : (0 : Finset α)⁻¹ = 0 := by ext; simp - -@[simp] lemma inv_smul_finset_distrib₀ (a : α) (s : Finset α) : (a • s)⁻¹ = op a⁻¹ • s⁻¹ := by - obtain rfl | ha := eq_or_ne a 0 - · obtain rfl | hs := s.eq_empty_or_nonempty <;> simp [*] - · ext; simp [← inv_smul_mem_iff₀, *] - -@[simp] lemma inv_op_smul_finset_distrib₀ (a : α) (s : Finset α) : (op a • s)⁻¹ = a⁻¹ • s⁻¹ := by - obtain rfl | ha := eq_or_ne a 0 - · obtain rfl | hs := s.eq_empty_or_nonempty <;> simp [*] - · ext; simp [← inv_smul_mem_iff₀, *] - -end GroupWithZero - -section Monoid - -variable [Monoid α] [AddGroup β] [DistribMulAction α β] [DecidableEq β] (a : α) (s : Finset α) - (t : Finset β) - -@[simp] -theorem smul_finset_neg : a • -t = -(a • t) := by - simp only [← image_smul, ← image_neg, Function.comp_def, image_image, smul_neg] - -@[simp] -protected theorem smul_neg : s • -t = -(s • t) := by - simp_rw [← image_neg] - exact image_image₂_right_comm smul_neg - -end Monoid - -section Ring - -variable [Ring α] [AddCommGroup β] [Module α β] [DecidableEq β] {s : Finset α} {t : Finset β} - {a : α} - -@[simp] -theorem neg_smul_finset : -a • t = -(a • t) := by - simp only [← image_smul, ← image_neg, image_image, neg_smul, Function.comp_def] - -@[simp] -protected theorem neg_smul [DecidableEq α] : -s • t = -(s • t) := by - simp_rw [← image_neg] - exact image₂_image_left_comm neg_smul - -end Ring - section BigOps section CommMonoid variable [CommMonoid α] {ι : Type*} [DecidableEq ι] @@ -2042,4 +1790,4 @@ instance Nat.decidablePred_mem_vadd_set {s : Set ℕ} [DecidablePred (· ∈ s)] fun n ↦ decidable_of_iff' (a ≤ n ∧ n - a ∈ s) <| by simp only [Set.mem_vadd_set, vadd_eq_add]; aesop -set_option linter.style.longFile 2100 +set_option linter.style.longFile 1800 diff --git a/Mathlib/Algebra/GroupWithZero/Pointwise/Finset.lean b/Mathlib/Algebra/GroupWithZero/Pointwise/Finset.lean new file mode 100644 index 0000000000000..7d763b441f0c9 --- /dev/null +++ b/Mathlib/Algebra/GroupWithZero/Pointwise/Finset.lean @@ -0,0 +1,209 @@ +/- +Copyright (c) 2021 Yaël Dillies, Bhavik Mehta. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies, Bhavik Mehta +-/ +import Mathlib.Algebra.Group.Pointwise.Finset.Basic + +/-! +# Pointwise operations of finsets in a group with zero + +This file proves properties of pointwise operations of finsets in a group with zero. +-/ + +-- TODO +-- assert_not_exists Ring + +open scoped Pointwise + +namespace Finset +variable {α β : Type*} [DecidableEq β] + +/-- If scalar multiplication by elements of `α` sends `(0 : β)` to zero, +then the same is true for `(0 : Finset β)`. -/ +protected def smulZeroClass [Zero β] [SMulZeroClass α β] : SMulZeroClass α (Finset β) := + coe_injective.smulZeroClass ⟨(↑), coe_zero⟩ coe_smul_finset + +/-- If the scalar multiplication `(· • ·) : α → β → β` is distributive, +then so is `(· • ·) : α → Finset β → Finset β`. -/ +protected def distribSMul [AddZeroClass β] [DistribSMul α β] : DistribSMul α (Finset β) := + coe_injective.distribSMul coeAddMonoidHom coe_smul_finset + +/-- A distributive multiplicative action of a monoid on an additive monoid `β` gives a distributive +multiplicative action on `Finset β`. -/ +protected def distribMulAction [Monoid α] [AddMonoid β] [DistribMulAction α β] : + DistribMulAction α (Finset β) := + coe_injective.distribMulAction coeAddMonoidHom coe_smul_finset + +/-- A multiplicative action of a monoid on a monoid `β` gives a multiplicative action on `Set β`. -/ +protected def mulDistribMulAction [Monoid α] [Monoid β] [MulDistribMulAction α β] : + MulDistribMulAction α (Finset β) := + coe_injective.mulDistribMulAction coeMonoidHom coe_smul_finset + +scoped[Pointwise] attribute [instance] Finset.smulZeroClass Finset.distribSMul + Finset.distribMulAction Finset.mulDistribMulAction + +instance [DecidableEq α] [Zero α] [Mul α] [NoZeroDivisors α] : NoZeroDivisors (Finset α) := + Function.Injective.noZeroDivisors (↑) coe_injective coe_zero coe_mul + +instance noZeroSMulDivisors [Zero α] [Zero β] [SMul α β] [NoZeroSMulDivisors α β] : + NoZeroSMulDivisors (Finset α) (Finset β) where + eq_zero_or_eq_zero_of_smul_eq_zero {s t} := by + exact_mod_cast eq_zero_or_eq_zero_of_smul_eq_zero (c := s.toSet) (x := t.toSet) + +instance noZeroSMulDivisors_finset [Zero α] [Zero β] [SMul α β] [NoZeroSMulDivisors α β] : + NoZeroSMulDivisors α (Finset β) := + Function.Injective.noZeroSMulDivisors (↑) coe_injective coe_zero coe_smul_finset + +section SMulZeroClass +variable [Zero β] [SMulZeroClass α β] {s : Finset α} {t : Finset β} {a : α} + +lemma smul_zero_subset (s : Finset α) : s • (0 : Finset β) ⊆ 0 := by simp [subset_iff, mem_smul] + +lemma Nonempty.smul_zero (hs : s.Nonempty) : s • (0 : Finset β) = 0 := + s.smul_zero_subset.antisymm <| by simpa [mem_smul] using hs + +lemma zero_mem_smul_finset (h : (0 : β) ∈ t) : (0 : β) ∈ a • t := + mem_smul_finset.2 ⟨0, h, smul_zero _⟩ + +variable [Zero α] [NoZeroSMulDivisors α β] + +lemma zero_mem_smul_finset_iff (ha : a ≠ 0) : (0 : β) ∈ a • t ↔ (0 : β) ∈ t := by + rw [← mem_coe, coe_smul_finset, Set.zero_mem_smul_set_iff ha, mem_coe] + +end SMulZeroClass + +section SMulWithZero +variable [Zero α] [Zero β] [SMulWithZero α β] {s : Finset α} {t : Finset β} + +/-! +Note that we have neither `SMulWithZero α (Finset β)` nor `SMulWithZero (Finset α) (Finset β)` +because `0 • ∅ ≠ 0`. +-/ + +lemma zero_smul_subset (t : Finset β) : (0 : Finset α) • t ⊆ 0 := by simp [subset_iff, mem_smul] + +lemma Nonempty.zero_smul (ht : t.Nonempty) : (0 : Finset α) • t = 0 := + t.zero_smul_subset.antisymm <| by simpa [mem_smul] using ht + +/-- A nonempty set is scaled by zero to the singleton set containing zero. -/ +@[simp] lemma zero_smul_finset {s : Finset β} (h : s.Nonempty) : (0 : α) • s = (0 : Finset β) := + coe_injective <| by simpa using @Set.zero_smul_set α _ _ _ _ _ h + +lemma zero_smul_finset_subset (s : Finset β) : (0 : α) • s ⊆ 0 := + image_subset_iff.2 fun x _ ↦ mem_zero.2 <| zero_smul α x + +variable [NoZeroSMulDivisors α β] + +lemma zero_mem_smul_iff : + (0 : β) ∈ s • t ↔ (0 : α) ∈ s ∧ t.Nonempty ∨ (0 : β) ∈ t ∧ s.Nonempty := by + rw [← mem_coe, coe_smul, Set.zero_mem_smul_iff]; rfl + +end SMulWithZero + +section MulZeroClass +variable [DecidableEq α] [MulZeroClass α] {s : Finset α} + +/-! Note that `Finset` is not a `MulZeroClass` because `0 * ∅ ≠ 0`. -/ + +lemma mul_zero_subset (s : Finset α) : s * 0 ⊆ 0 := by simp [subset_iff, mem_mul] +lemma zero_mul_subset (s : Finset α) : 0 * s ⊆ 0 := by simp [subset_iff, mem_mul] + +lemma Nonempty.mul_zero (hs : s.Nonempty) : s * 0 = 0 := + s.mul_zero_subset.antisymm <| by simpa [mem_mul] using hs + +lemma Nonempty.zero_mul (hs : s.Nonempty) : 0 * s = 0 := + s.zero_mul_subset.antisymm <| by simpa [mem_mul] using hs + +end MulZeroClass + +section GroupWithZero +variable [GroupWithZero α] + +section MulAction +variable [MulAction α β] {s t : Finset β} {a : α} {b : β} + +@[simp] lemma smul_mem_smul_finset_iff₀ (ha : a ≠ 0) : a • b ∈ a • s ↔ b ∈ s := + smul_mem_smul_finset_iff (Units.mk0 a ha) + +lemma inv_smul_mem_iff₀ (ha : a ≠ 0) : a⁻¹ • b ∈ s ↔ b ∈ a • s := + show _ ↔ _ ∈ Units.mk0 a ha • _ from inv_smul_mem_iff + +lemma mem_inv_smul_finset_iff₀ (ha : a ≠ 0) : b ∈ a⁻¹ • s ↔ a • b ∈ s := + show _ ∈ (Units.mk0 a ha)⁻¹ • _ ↔ _ from mem_inv_smul_finset_iff + +@[simp] +lemma smul_finset_subset_smul_finset_iff₀ (ha : a ≠ 0) : a • s ⊆ a • t ↔ s ⊆ t := + show Units.mk0 a ha • _ ⊆ _ ↔ _ from smul_finset_subset_smul_finset_iff + +lemma smul_finset_subset_iff₀ (ha : a ≠ 0) : a • s ⊆ t ↔ s ⊆ a⁻¹ • t := + show Units.mk0 a ha • _ ⊆ _ ↔ _ from smul_finset_subset_iff + +lemma subset_smul_finset_iff₀ (ha : a ≠ 0) : s ⊆ a • t ↔ a⁻¹ • s ⊆ t := + show _ ⊆ Units.mk0 a ha • _ ↔ _ from subset_smul_finset_iff + +lemma smul_finset_inter₀ (ha : a ≠ 0) : a • (s ∩ t) = a • s ∩ a • t := + image_inter _ _ <| MulAction.injective₀ ha + +lemma smul_finset_sdiff₀ (ha : a ≠ 0) : a • (s \ t) = a • s \ a • t := + image_sdiff _ _ <| MulAction.injective₀ ha + +open scoped symmDiff in +lemma smul_finset_symmDiff₀ (ha : a ≠ 0) : a • s ∆ t = (a • s) ∆ (a • t) := + image_symmDiff _ _ <| MulAction.injective₀ ha + +lemma smul_finset_univ₀ [Fintype β] (ha : a ≠ 0) : a • (univ : Finset β) = univ := + coe_injective <| by push_cast; exact Set.smul_set_univ₀ ha + +lemma smul_univ₀ [Fintype β] {s : Finset α} (hs : ¬s ⊆ 0) : s • (univ : Finset β) = univ := + coe_injective <| by + rw [← coe_subset] at hs + push_cast at hs ⊢ + exact Set.smul_univ₀ hs + +lemma smul_univ₀' [Fintype β] {s : Finset α} (hs : s.Nontrivial) : s • (univ : Finset β) = univ := + coe_injective <| by push_cast; exact Set.smul_univ₀' hs + +end MulAction + +variable [DecidableEq α] {s : Finset α} + +lemma div_zero_subset (s : Finset α) : s / 0 ⊆ 0 := by simp [subset_iff, mem_div] + +lemma zero_div_subset (s : Finset α) : 0 / s ⊆ 0 := by simp [subset_iff, mem_div] + +lemma Nonempty.div_zero (hs : s.Nonempty) : s / 0 = 0 := + s.div_zero_subset.antisymm <| by simpa [mem_div] using hs + +lemma Nonempty.zero_div (hs : s.Nonempty) : 0 / s = 0 := + s.zero_div_subset.antisymm <| by simpa [mem_div] using hs + +@[simp] protected lemma inv_zero : (0 : Finset α)⁻¹ = 0 := by ext; simp + +open scoped RightActions + +@[simp] lemma inv_smul_finset_distrib₀ (a : α) (s : Finset α) : (a • s)⁻¹ = s⁻¹ <• a⁻¹ := by + obtain rfl | ha := eq_or_ne a 0 + · obtain rfl | hs := s.eq_empty_or_nonempty <;> simp [*] + · ext; simp [← inv_smul_mem_iff₀, *] + +@[simp] lemma inv_op_smul_finset_distrib₀ (a : α) (s : Finset α) : (s <• a)⁻¹ = a⁻¹ • s⁻¹ := by + obtain rfl | ha := eq_or_ne a 0 + · obtain rfl | hs := s.eq_empty_or_nonempty <;> simp [*] + · ext; simp [← inv_smul_mem_iff₀, *] + +end GroupWithZero + +section Monoid +variable [Monoid α] [AddGroup β] [DistribMulAction α β] + +@[simp] +lemma smul_finset_neg (a : α) (t : Finset β) : a • -t = -(a • t) := by + simp only [← image_smul, ← image_neg, Function.comp_def, image_image, smul_neg] + +@[simp] +protected lemma smul_neg (s : Finset α) (t : Finset β) : s • -t = -(s • t) := by + simp_rw [← image_neg]; exact image_image₂_right_comm smul_neg + +end Monoid +end Finset diff --git a/Mathlib/Algebra/Ring/Pointwise/Finset.lean b/Mathlib/Algebra/Ring/Pointwise/Finset.lean new file mode 100644 index 0000000000000..bd2a3b4e8a7f9 --- /dev/null +++ b/Mathlib/Algebra/Ring/Pointwise/Finset.lean @@ -0,0 +1,69 @@ +/- +Copyright (c) 2022 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies +-/ +import Mathlib.Algebra.Group.Pointwise.Finset.Basic +import Mathlib.Algebra.Ring.Pointwise.Set + +/-! +# Pointwise operations of sets in a ring + +This file proves properties of pointwise operations of sets in a ring. + +## Tags + +set multiplication, set addition, pointwise addition, pointwise multiplication, +pointwise subtraction +-/ + +open scoped Pointwise + +namespace Finset +variable {α β : Type*} + +/-- `Finset α` has distributive negation if `α` has. -/ +protected def distribNeg [DecidableEq α] [Mul α] [HasDistribNeg α] : HasDistribNeg (Finset α) := + coe_injective.hasDistribNeg _ coe_neg coe_mul + +scoped[Pointwise] attribute [instance] Finset.distribNeg + +section Distrib +variable [DecidableEq α] [Distrib α] (s t u : Finset α) + +/-! +Note that `Finset α` is not a `Distrib` because `s * t + s * u` has cross terms that `s * (t + u)` +lacks. + +```lean +-- {10, 16, 18, 20, 8, 9} +#eval {1, 2} * ({3, 4} + {5, 6} : Finset ℕ) + +-- {10, 11, 12, 13, 14, 15, 16, 18, 20, 8, 9} +#eval ({1, 2} : Finset ℕ) * {3, 4} + {1, 2} * {5, 6} +``` +-/ + +lemma mul_add_subset : s * (t + u) ⊆ s * t + s * u := + image₂_distrib_subset_left mul_add + +lemma add_mul_subset : (s + t) * u ⊆ s * u + t * u := + image₂_distrib_subset_right add_mul + +end Distrib + +section Ring +variable [Ring α] [AddCommGroup β] [Module α β] [DecidableEq β] {s : Finset α} {t : Finset β} + {a : α} + +@[simp] +lemma neg_smul_finset : -a • t = -(a • t) := by + simp only [← image_smul, ← image_neg, image_image, neg_smul, Function.comp_def] + +@[simp] +protected lemma neg_smul [DecidableEq α] : -s • t = -(s • t) := by + simp_rw [← image_neg] + exact image₂_image_left_comm neg_smul + +end Ring +end Finset From ff8e6e85cd0dc2dd27af8d21c731a04ce74b943e Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Mon, 21 Oct 2024 16:43:40 +0000 Subject: [PATCH 246/505] feat: semisimple rings are Artinian (#17557) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The main result `IsSemisimpleModule.finite_tfae` shows that for a semisimple module, `Module.Finite R M ↔ IsNoetherian R M ↔ IsArtinian R M ↔ IsFiniteLength R M ↔` direct sum decomposition into finitely many simple modules (F25 in Chapter 28 of Lorenz's *Algebra II*). The instance from IsSemisimpleModule + Module.Finite to IsNoetherian is proven before and used in the proof, while the instance to IsArtinian is added after the proof. As a consequence, semisimple rings are Artinian. Also add easy lemmas `isNoetherian_range`, `LinearEquiv.isNoetherian_iff`, `isNoetherian_sup`, `isNoetherian_iSup` (and Artinian counterparts) and `IsSemisimpleModule.exists_setIndependent_sSup_simples_eq_top`. --- Mathlib/Algebra/DirectSum/Module.lean | 5 ++- .../NumberField/Discriminant.lean | 4 +-- Mathlib/RingTheory/Artinian.lean | 33 +++++++++++++----- Mathlib/RingTheory/FiniteLength.lean | 28 ++++++++++++++- Mathlib/RingTheory/Noetherian.lean | 34 ++++++++++++++----- Mathlib/RingTheory/SimpleModule.lean | 14 ++++++++ test/instances/Ring_finiteness.lean | 8 +++++ 7 files changed, 105 insertions(+), 21 deletions(-) create mode 100644 test/instances/Ring_finiteness.lean diff --git a/Mathlib/Algebra/DirectSum/Module.lean b/Mathlib/Algebra/DirectSum/Module.lean index 4ed1b6c1db1b8..0b82afa6bcc70 100644 --- a/Mathlib/Algebra/DirectSum/Module.lean +++ b/Mathlib/Algebra/DirectSum/Module.lean @@ -269,7 +269,7 @@ variable {ι : Type v} [dec_ι : DecidableEq ι] variable {M : Type*} [AddCommMonoid M] [Module R M] variable (A : ι → Submodule R M) -/-- The canonical embedding from `⨁ i, A i` to `M` where `A` is a collection of `Submodule R M` +/-- The canonical linear map from `⨁ i, A i` to `M` where `A` is a collection of `Submodule R M` indexed by `ι`. This is `DirectSum.coeAddMonoidHom` as a `LinearMap`. -/ def coeLinearMap : (⨁ i, A i) →ₗ[R] M := toModule R ι M fun i ↦ (A i).subtype @@ -288,6 +288,9 @@ theorem coeLinearMap_of (i : ι) (x : A i) : DirectSum.coeLinearMap A (of (fun i variable {A} +theorem range_coeLinearMap : LinearMap.range (coeLinearMap A) = ⨆ i, A i := + (Submodule.iSup_eq_range_dfinsupp_lsum _).symm + @[simp] theorem IsInternal.ofBijective_coeLinearMap_same (h : IsInternal A) {i : ι} (x : A i) : diff --git a/Mathlib/NumberTheory/NumberField/Discriminant.lean b/Mathlib/NumberTheory/NumberField/Discriminant.lean index 07199e222dcb4..393c5e32016a0 100644 --- a/Mathlib/NumberTheory/NumberField/Discriminant.lean +++ b/Mathlib/NumberTheory/NumberField/Discriminant.lean @@ -360,7 +360,7 @@ theorem finite_of_discr_bdd_of_isReal : · exact (Nat.choose_le_choose _ (rank_le_rankOfDiscrBdd hK₂)).trans (Nat.choose_le_middle _ _) · refine mem_rootSet.mpr ⟨minpoly.ne_zero hx, ?_⟩ - exact (aeval_algebraMap_eq_zero_iff _ _ _).mpr (minpoly.aeval ℤ (x : K)) + exact (aeval_algebraMap_eq_zero_iff A (x : K) _).mpr (minpoly.aeval ℤ (x : K)) · rw [← (IntermediateField.lift_injective _).eq_iff, eq_comm] at hx₁ convert hx₁ · simp only [IntermediateField.lift_top] @@ -409,7 +409,7 @@ theorem finite_of_discr_bdd_of_isComplex : exact (Nat.choose_le_choose _ (rank_le_rankOfDiscrBdd hK₂)).trans (Nat.choose_le_middle _ _) · refine mem_rootSet.mpr ⟨minpoly.ne_zero hx, ?_⟩ - exact (aeval_algebraMap_eq_zero_iff _ _ _).mpr (minpoly.aeval ℤ (x : K)) + exact (aeval_algebraMap_eq_zero_iff A (x : K) _).mpr (minpoly.aeval ℤ (x : K)) · rw [← (IntermediateField.lift_injective _).eq_iff, eq_comm] at hx₁ convert hx₁ · simp only [IntermediateField.lift_top] diff --git a/Mathlib/RingTheory/Artinian.lean b/Mathlib/RingTheory/Artinian.lean index 9e43f7cbc62b9..6942b8afeb66c 100644 --- a/Mathlib/RingTheory/Artinian.lean +++ b/Mathlib/RingTheory/Artinian.lean @@ -97,9 +97,15 @@ instance isArtinian_of_quotient_of_artinian variable {M} +instance isArtinian_range (f : M →ₗ[R] P) [IsArtinian R M] : IsArtinian R (LinearMap.range f) := + isArtinian_of_surjective _ _ f.surjective_rangeRestrict + theorem isArtinian_of_linearEquiv (f : M ≃ₗ[R] P) [IsArtinian R M] : IsArtinian R P := isArtinian_of_surjective _ f.toLinearMap f.toEquiv.surjective +theorem LinearEquiv.isArtinian_iff (f : M ≃ₗ[R] P) : IsArtinian R M ↔ IsArtinian R P := + ⟨fun _ ↦ isArtinian_of_linearEquiv f, fun _ ↦ isArtinian_of_linearEquiv f.symm⟩ + theorem isArtinian_of_range_eq_ker [IsArtinian R M] [IsArtinian R P] (f : M →ₗ[R] N) (g : N →ₗ[R] P) (h : LinearMap.range f = LinearMap.ker g) : IsArtinian R N := wellFounded_lt_exact_sequence (LinearMap.range f) (Submodule.map (f.ker.liftQ f le_rfl)) @@ -125,8 +131,15 @@ instance (priority := 100) isArtinian_of_finite [Finite M] : IsArtinian R M := -- Porting note: elab_as_elim can only be global and cannot be changed on an imported decl -- attribute [local elab_as_elim] Finite.induction_empty_option -instance isArtinian_pi {R ι : Type*} [Finite ι] : - ∀ {M : ι → Type*} [Ring R] [∀ i, AddCommGroup (M i)] +instance isArtinian_sup (M₁ M₂ : Submodule R P) [IsArtinian R M₁] [IsArtinian R M₂] : + IsArtinian R ↥(M₁ ⊔ M₂) := by + have := isArtinian_range (M₁.subtype.coprod M₂.subtype) + rwa [LinearMap.range_coprod, Submodule.range_subtype, Submodule.range_subtype] at this + +variable {ι : Type*} [Finite ι] + +instance isArtinian_pi : + ∀ {M : ι → Type*} [∀ i, AddCommGroup (M i)] [∀ i, Module R (M i)] [∀ i, IsArtinian R (M i)], IsArtinian R (∀ i, M i) := by apply Finite.induction_empty_option _ _ _ ι · exact fun e h ↦ isArtinian_of_linearEquiv (LinearEquiv.piCongrLeft R _ e) @@ -136,15 +149,20 @@ instance isArtinian_pi {R ι : Type*} [Finite ι] : /-- A version of `isArtinian_pi` for non-dependent functions. We need this instance because sometimes Lean fails to apply the dependent version in non-dependent settings (e.g., it fails to prove that `ι → ℝ` is finite dimensional over `ℝ`). -/ -instance isArtinian_pi' {R ι M : Type*} [Ring R] [AddCommGroup M] [Module R M] [Finite ι] - [IsArtinian R M] : IsArtinian R (ι → M) := +instance isArtinian_pi' [IsArtinian R M] : IsArtinian R (ι → M) := isArtinian_pi --porting note (#10754): new instance -instance isArtinian_finsupp {R ι M : Type*} [Ring R] [AddCommGroup M] [Module R M] [Finite ι] - [IsArtinian R M] : IsArtinian R (ι →₀ M) := +instance isArtinian_finsupp [IsArtinian R M] : IsArtinian R (ι →₀ M) := isArtinian_of_linearEquiv (Finsupp.linearEquivFunOnFinite _ _ _).symm +instance isArtinian_iSup : + ∀ {M : ι → Submodule R P} [∀ i, IsArtinian R (M i)], IsArtinian R ↥(⨆ i, M i) := by + apply Finite.induction_empty_option _ _ _ ι + · intro _ _ e h _ _; rw [← e.iSup_comp]; apply h + · intros; rw [iSup_of_empty]; infer_instance + · intro _ _ ih _ _; rw [iSup_option]; infer_instance + end open Submodule Function @@ -522,7 +540,4 @@ instance [IsReduced R] : DecompositionMonoid (Polynomial R) := theorem isSemisimpleRing_of_isReduced [IsReduced R] : IsSemisimpleRing R := (equivPi R).symm.isSemisimpleRing -proof_wanted IsSemisimpleRing.isArtinianRing (R : Type*) [Ring R] [IsSemisimpleRing R] : - IsArtinianRing R - end IsArtinianRing diff --git a/Mathlib/RingTheory/FiniteLength.lean b/Mathlib/RingTheory/FiniteLength.lean index d0305806bd193..b1cf98eb550b8 100644 --- a/Mathlib/RingTheory/FiniteLength.lean +++ b/Mathlib/RingTheory/FiniteLength.lean @@ -3,7 +3,6 @@ Copyright (c) 2024 Junyan Xu. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Junyan Xu -/ -import Mathlib.Order.Atoms.Finite import Mathlib.RingTheory.Artinian /-! @@ -76,3 +75,30 @@ theorem isFiniteLength_iff_exists_compositionSeries : ⟨fun h ↦ have ⟨_, _⟩ := isFiniteLength_iff_isNoetherian_isArtinian.mp h exists_compositionSeries_of_isNoetherian_isArtinian R M, isFiniteLength_of_exists_compositionSeries⟩ + +theorem IsSemisimpleModule.finite_tfae [IsSemisimpleModule R M] : + List.TFAE [Module.Finite R M, IsNoetherian R M, IsArtinian R M, IsFiniteLength R M, + ∃ s : Set (Submodule R M), s.Finite ∧ CompleteLattice.SetIndependent s ∧ + sSup s = ⊤ ∧ ∀ m ∈ s, IsSimpleModule R m] := by + rw [isFiniteLength_iff_isNoetherian_isArtinian] + obtain ⟨s, hs⟩ := IsSemisimpleModule.exists_setIndependent_sSup_simples_eq_top R M + tfae_have 1 ↔ 2 := ⟨fun _ ↦ inferInstance, fun _ ↦ inferInstance⟩ + tfae_have 2 → 5 := fun _ ↦ ⟨s, CompleteLattice.WellFoundedGT.finite_of_setIndependent hs.1, hs⟩ + tfae_have 3 → 5 := fun _ ↦ ⟨s, CompleteLattice.WellFoundedLT.finite_of_setIndependent hs.1, hs⟩ + tfae_have 5 → 4 := fun ⟨s, fin, _, sSup_eq_top, simple⟩ ↦ by + rw [← isNoetherian_top_iff, ← Submodule.topEquiv.isArtinian_iff, + ← sSup_eq_top, sSup_eq_iSup, ← iSup_subtype''] + rw [SetCoe.forall'] at simple + have := fin.to_subtype + exact ⟨isNoetherian_iSup, isArtinian_iSup⟩ + tfae_have 4 → 2 := And.left + tfae_have 4 → 3 := And.right + tfae_finish + +instance [IsSemisimpleModule R M] [Module.Finite R M] : IsArtinian R M := + (IsSemisimpleModule.finite_tfae.out 0 2).mp ‹_› + +/- The following instances are now automatic: +example [IsSemisimpleRing R] : IsNoetherianRing R := inferInstance +example [IsSemisimpleRing R] : IsArtinianRing R := inferInstance +-/ diff --git a/Mathlib/RingTheory/Noetherian.lean b/Mathlib/RingTheory/Noetherian.lean index 411b9fdf6c5a4..54ede4abbaac0 100644 --- a/Mathlib/RingTheory/Noetherian.lean +++ b/Mathlib/RingTheory/Noetherian.lean @@ -112,6 +112,10 @@ theorem isNoetherian_of_surjective (f : M →ₗ[R] P) (hf : LinearMap.range f = variable {M} +instance isNoetherian_range (f : M →ₗ[R] P) [IsNoetherian R M] : + IsNoetherian R (LinearMap.range f) := + isNoetherian_of_surjective _ _ f.range_rangeRestrict + instance isNoetherian_quotient {A M : Type*} [Ring A] [AddCommGroup M] [SMul R A] [Module R M] [Module A M] [IsScalarTower R A M] (N : Submodule A M) [IsNoetherian R M] : IsNoetherian R (M ⧸ N) := @@ -124,10 +128,11 @@ alias Submodule.Quotient.isNoetherian := isNoetherian_quotient theorem isNoetherian_of_linearEquiv (f : M ≃ₗ[R] P) [IsNoetherian R M] : IsNoetherian R P := isNoetherian_of_surjective _ f.toLinearMap f.range -theorem isNoetherian_top_iff : IsNoetherian R (⊤ : Submodule R M) ↔ IsNoetherian R M := by - constructor <;> intro h - · exact isNoetherian_of_linearEquiv (LinearEquiv.ofTop (⊤ : Submodule R M) rfl) - · exact isNoetherian_of_linearEquiv (LinearEquiv.ofTop (⊤ : Submodule R M) rfl).symm +theorem LinearEquiv.isNoetherian_iff (f : M ≃ₗ[R] P) : IsNoetherian R M ↔ IsNoetherian R P := + ⟨fun _ ↦ isNoetherian_of_linearEquiv f, fun _ ↦ isNoetherian_of_linearEquiv f.symm⟩ + +theorem isNoetherian_top_iff : IsNoetherian R (⊤ : Submodule R M) ↔ IsNoetherian R M := + Submodule.topEquiv.isNoetherian_iff theorem isNoetherian_of_injective [IsNoetherian R P] (f : M →ₗ[R] P) (hf : Function.Injective f) : IsNoetherian R M := @@ -190,8 +195,15 @@ 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*} [Finite ι] : - ∀ {M : ι → Type*} [Ring R] [∀ i, AddCommGroup (M i)] +instance isNoetherian_sup (M₁ M₂ : Submodule R P) [IsNoetherian R M₁] [IsNoetherian R M₂] : + IsNoetherian R ↥(M₁ ⊔ M₂) := by + have := isNoetherian_range (M₁.subtype.coprod M₂.subtype) + rwa [LinearMap.range_coprod, Submodule.range_subtype, Submodule.range_subtype] at this + +variable {ι : Type*} [Finite ι] + +instance isNoetherian_pi : + ∀ {M : ι → Type*} [∀ i, AddCommGroup (M i)] [∀ i, Module R (M i)] [∀ i, IsNoetherian R (M i)], IsNoetherian R (∀ i, M i) := by apply Finite.induction_empty_option _ _ _ ι · exact fun e h ↦ isNoetherian_of_linearEquiv (LinearEquiv.piCongrLeft R _ e) @@ -201,10 +213,16 @@ instance isNoetherian_pi {R ι : Type*} [Finite ι] : /-- 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 prove that `ι → ℝ` is finite dimensional over `ℝ`). -/ -instance isNoetherian_pi' {R ι M : Type*} [Ring R] [AddCommGroup M] [Module R M] [Finite ι] - [IsNoetherian R M] : IsNoetherian R (ι → M) := +instance isNoetherian_pi' [IsNoetherian R M] : IsNoetherian R (ι → M) := isNoetherian_pi +instance isNoetherian_iSup : + ∀ {M : ι → Submodule R P} [∀ i, IsNoetherian R (M i)], IsNoetherian R ↥(⨆ i, M i) := by + apply Finite.induction_empty_option _ _ _ ι + · intro _ _ e h _ _; rw [← e.iSup_comp]; apply h + · intros; rw [iSup_of_empty]; infer_instance + · intro _ _ ih _ _; rw [iSup_option]; infer_instance + end section CommRing diff --git a/Mathlib/RingTheory/SimpleModule.lean b/Mathlib/RingTheory/SimpleModule.lean index ca9632f9632b5..1eddfcca1ca75 100644 --- a/Mathlib/RingTheory/SimpleModule.lean +++ b/Mathlib/RingTheory/SimpleModule.lean @@ -5,6 +5,7 @@ Authors: Aaron Anderson -/ import Mathlib.LinearAlgebra.Isomorphisms import Mathlib.LinearAlgebra.Projection +import Mathlib.Order.Atoms.Finite import Mathlib.Order.JordanHolder import Mathlib.Order.CompactlyGenerated.Intervals import Mathlib.LinearAlgebra.FiniteDimensional @@ -123,6 +124,8 @@ theorem ker_toSpanSingleton_isMaximal {m : M} (hm : m ≠ 0) : rw [Ideal.isMaximal_def, ← isSimpleModule_iff_isCoatom] exact congr (quotKerEquivOfSurjective _ <| toSpanSingleton_surjective R hm) +instance : IsNoetherian R M := isNoetherian_iff'.mpr inferInstance + end IsSimpleModule open IsSimpleModule in @@ -194,6 +197,13 @@ theorem exists_simple_submodule [Nontrivial M] : ∃ m : Submodule R M, IsSimple theorem sSup_simples_eq_top : sSup { m : Submodule R M | IsSimpleModule R m } = ⊤ := by simpa only [isSimpleModule_iff_isAtom] using sSup_atoms_eq_top +theorem exists_setIndependent_sSup_simples_eq_top : + ∃ s : Set (Submodule R M), CompleteLattice.SetIndependent s ∧ + sSup s = ⊤ ∧ ∀ m ∈ s, IsSimpleModule R m := by + have := sSup_simples_eq_top R M + simp_rw [isSimpleModule_iff_isAtom] at this ⊢ + exact exists_setIndependent_of_sSup_atoms_eq_top this + /-- The annihilator of a semisimple module over a commutative ring is a radical ideal. -/ theorem annihilator_isRadical (R) [CommRing R] [Module R M] [IsSemisimpleModule R M] : (Module.annihilator R M).IsRadical := by @@ -213,6 +223,10 @@ instance quotient : IsSemisimpleModule R (M ⧸ m) := have ⟨P, compl⟩ := exists_isCompl m .congr (m.quotientEquivOfIsCompl P compl) +instance (priority := low) [Module.Finite R M] : IsNoetherian R M where + noetherian m := have ⟨P, compl⟩ := exists_isCompl m + Module.Finite.iff_fg.mp (Module.Finite.equiv <| P.quotientEquivOfIsCompl m compl.symm) + -- does not work as an instance, not sure why protected theorem range (f : M →ₗ[R] N) : IsSemisimpleModule R (range f) := .congr (quotKerEquivRange _).symm diff --git a/test/instances/Ring_finiteness.lean b/test/instances/Ring_finiteness.lean new file mode 100644 index 0000000000000..89e9e0a99e860 --- /dev/null +++ b/test/instances/Ring_finiteness.lean @@ -0,0 +1,8 @@ +import Mathlib + +-- https://github.com/leanprover-community/mathlib4/pull/17557#issuecomment-2426920648 + +variable (R : Type*) [Ring R] [IsSemisimpleRing R] + +example : IsNoetherianRing R := inferInstance +example : IsArtinianRing R := inferInstance From d2c8beddb0e73ffe98e51d3a9933d5acd833d52f Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 21 Oct 2024 16:43:41 +0000 Subject: [PATCH 247/505] chore(Topology/Algebra/SeparationQuotient): split (#18011) The last 50 lines of the file cause massive import increase in downstream files. --- Mathlib.lean | 3 +- Mathlib/Analysis/LocallyConvex/Bounded.lean | 1 + Mathlib/Analysis/Normed/Module/Span.lean | 1 + .../Normed/Operator/ContinuousLinearMap.lean | 1 + .../Normed/Operator/LinearIsometry.lean | 3 +- Mathlib/Analysis/NormedSpace/BallAction.lean | 1 + .../NormedSpace/ConformalLinearMap.lean | 1 + Mathlib/Analysis/NormedSpace/ENorm.lean | 1 + Mathlib/Analysis/NormedSpace/Real.lean | 1 + .../Analysis/NormedSpace/SphereNormEquiv.lean | 1 + Mathlib/Analysis/RCLike/Basic.lean | 1 + .../Topology/Algebra/Module/Cardinality.lean | 5 +- .../Basic.lean} | 80 +--------------- .../Algebra/SeparationQuotient/Section.lean | 93 +++++++++++++++++++ Mathlib/Topology/MetricSpace/Algebra.lean | 2 +- Mathlib/Topology/PartitionOfUnity.lean | 1 + 16 files changed, 113 insertions(+), 83 deletions(-) rename Mathlib/Topology/Algebra/{SeparationQuotient.lean => SeparationQuotient/Basic.lean} (83%) create mode 100644 Mathlib/Topology/Algebra/SeparationQuotient/Section.lean diff --git a/Mathlib.lean b/Mathlib.lean index 539841cfd97d9..9ca592ec191d3 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4600,7 +4600,8 @@ import Mathlib.Topology.Algebra.ProperConstSMul import Mathlib.Topology.Algebra.Ring.Basic import Mathlib.Topology.Algebra.Ring.Ideal import Mathlib.Topology.Algebra.Semigroup -import Mathlib.Topology.Algebra.SeparationQuotient +import Mathlib.Topology.Algebra.SeparationQuotient.Basic +import Mathlib.Topology.Algebra.SeparationQuotient.Section import Mathlib.Topology.Algebra.Star import Mathlib.Topology.Algebra.StarSubalgebra import Mathlib.Topology.Algebra.UniformConvergence diff --git a/Mathlib/Analysis/LocallyConvex/Bounded.lean b/Mathlib/Analysis/LocallyConvex/Bounded.lean index 8ec28cb68b8ab..e345035096ce2 100644 --- a/Mathlib/Analysis/LocallyConvex/Bounded.lean +++ b/Mathlib/Analysis/LocallyConvex/Bounded.lean @@ -7,6 +7,7 @@ import Mathlib.GroupTheory.GroupAction.Pointwise import Mathlib.Analysis.LocallyConvex.Basic import Mathlib.Analysis.LocallyConvex.BalancedCoreHull import Mathlib.Analysis.Seminorm +import Mathlib.LinearAlgebra.Basis.VectorSpace import Mathlib.Topology.Bornology.Basic import Mathlib.Topology.Algebra.UniformGroup import Mathlib.Topology.UniformSpace.Cauchy diff --git a/Mathlib/Analysis/Normed/Module/Span.lean b/Mathlib/Analysis/Normed/Module/Span.lean index 2996a80f34809..cbb4571a230b6 100644 --- a/Mathlib/Analysis/Normed/Module/Span.lean +++ b/Mathlib/Analysis/Normed/Module/Span.lean @@ -7,6 +7,7 @@ Authors: Moritz Doll import Mathlib.Analysis.Normed.Operator.LinearIsometry import Mathlib.Analysis.Normed.Operator.ContinuousLinearMap import Mathlib.Analysis.Normed.Module.Basic +import Mathlib.LinearAlgebra.Basis.VectorSpace /-! # The span of a single vector diff --git a/Mathlib/Analysis/Normed/Operator/ContinuousLinearMap.lean b/Mathlib/Analysis/Normed/Operator/ContinuousLinearMap.lean index 64c5ba0eb60b5..8b61db66f18eb 100644 --- a/Mathlib/Analysis/Normed/Operator/ContinuousLinearMap.lean +++ b/Mathlib/Analysis/Normed/Operator/ContinuousLinearMap.lean @@ -5,6 +5,7 @@ Authors: Jan-David Salchow, Sébastien Gouëzel, Jean Lo -/ import Mathlib.Analysis.Normed.Group.Uniform import Mathlib.Analysis.Normed.MulAction +import Mathlib.LinearAlgebra.DFinsupp import Mathlib.Topology.Algebra.Module.Basic /-! # Constructions of continuous linear maps between (semi-)normed spaces diff --git a/Mathlib/Analysis/Normed/Operator/LinearIsometry.lean b/Mathlib/Analysis/Normed/Operator/LinearIsometry.lean index 92b76035981f0..d3192d513de52 100644 --- a/Mathlib/Analysis/Normed/Operator/LinearIsometry.lean +++ b/Mathlib/Analysis/Normed/Operator/LinearIsometry.lean @@ -7,8 +7,9 @@ import Mathlib.Algebra.Star.Basic import Mathlib.Analysis.Normed.Group.Constructions import Mathlib.Analysis.Normed.Group.Submodule import Mathlib.Analysis.Normed.Group.Uniform -import Mathlib.Topology.Algebra.Module.Basic import Mathlib.LinearAlgebra.Basis.Defs +import Mathlib.LinearAlgebra.DFinsupp +import Mathlib.Topology.Algebra.Module.Basic /-! # (Semi-)linear isometries diff --git a/Mathlib/Analysis/NormedSpace/BallAction.lean b/Mathlib/Analysis/NormedSpace/BallAction.lean index f095bd6baf5d4..2f484459d1044 100644 --- a/Mathlib/Analysis/NormedSpace/BallAction.lean +++ b/Mathlib/Analysis/NormedSpace/BallAction.lean @@ -5,6 +5,7 @@ Authors: Yury Kudryashov, Heather Macbeth -/ import Mathlib.Analysis.Normed.Field.UnitBall import Mathlib.Analysis.Normed.Module.Basic +import Mathlib.LinearAlgebra.Basis.VectorSpace /-! # Multiplicative actions of/on balls and spheres diff --git a/Mathlib/Analysis/NormedSpace/ConformalLinearMap.lean b/Mathlib/Analysis/NormedSpace/ConformalLinearMap.lean index d73ac04e060f8..068091b76b338 100644 --- a/Mathlib/Analysis/NormedSpace/ConformalLinearMap.lean +++ b/Mathlib/Analysis/NormedSpace/ConformalLinearMap.lean @@ -5,6 +5,7 @@ Authors: Yourong Zang -/ import Mathlib.Analysis.Normed.Module.Basic import Mathlib.Analysis.Normed.Operator.LinearIsometry +import Mathlib.LinearAlgebra.Basis.VectorSpace /-! # Conformal Linear Maps diff --git a/Mathlib/Analysis/NormedSpace/ENorm.lean b/Mathlib/Analysis/NormedSpace/ENorm.lean index 313555960cd59..5f64698d8c8c7 100644 --- a/Mathlib/Analysis/NormedSpace/ENorm.lean +++ b/Mathlib/Analysis/NormedSpace/ENorm.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import Mathlib.Analysis.Normed.Module.Basic +import Mathlib.LinearAlgebra.Basis.VectorSpace /-! # Extended norm diff --git a/Mathlib/Analysis/NormedSpace/Real.lean b/Mathlib/Analysis/NormedSpace/Real.lean index 3106fe40f746a..09c0aa69993d6 100644 --- a/Mathlib/Analysis/NormedSpace/Real.lean +++ b/Mathlib/Analysis/NormedSpace/Real.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov, Patrick Massot, Eric Wieser, Yaël Dillies -/ import Mathlib.Analysis.Normed.Module.Basic +import Mathlib.LinearAlgebra.Basis.VectorSpace import Mathlib.Topology.Algebra.Module.Basic /-! diff --git a/Mathlib/Analysis/NormedSpace/SphereNormEquiv.lean b/Mathlib/Analysis/NormedSpace/SphereNormEquiv.lean index fd8a5748ac567..c0c9a90640f7c 100644 --- a/Mathlib/Analysis/NormedSpace/SphereNormEquiv.lean +++ b/Mathlib/Analysis/NormedSpace/SphereNormEquiv.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import Mathlib.Analysis.Normed.Module.Basic +import Mathlib.LinearAlgebra.Basis.VectorSpace /-! # Homeomorphism between a normed space and sphere times `(0, +∞)` diff --git a/Mathlib/Analysis/RCLike/Basic.lean b/Mathlib/Analysis/RCLike/Basic.lean index 7db8b7b8d5973..ad002383d932e 100644 --- a/Mathlib/Analysis/RCLike/Basic.lean +++ b/Mathlib/Analysis/RCLike/Basic.lean @@ -10,6 +10,7 @@ import Mathlib.Algebra.Order.Star.Basic import Mathlib.Analysis.CStarAlgebra.Basic import Mathlib.Analysis.Normed.Operator.ContinuousLinearMap import Mathlib.Data.Real.Sqrt +import Mathlib.LinearAlgebra.Basis.VectorSpace /-! # `RCLike`: a typeclass for ℝ or ℂ diff --git a/Mathlib/Topology/Algebra/Module/Cardinality.lean b/Mathlib/Topology/Algebra/Module/Cardinality.lean index b5c7c62b830cd..845e4610d726e 100644 --- a/Mathlib/Topology/Algebra/Module/Cardinality.lean +++ b/Mathlib/Topology/Algebra/Module/Cardinality.lean @@ -4,9 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ import Mathlib.Algebra.Module.Card -import Mathlib.SetTheory.Cardinal.CountableCover -import Mathlib.SetTheory.Cardinal.Continuum import Mathlib.Analysis.SpecificLimits.Normed +import Mathlib.SetTheory.Cardinal.Continuum +import Mathlib.SetTheory.Cardinal.CountableCover +import Mathlib.LinearAlgebra.Basis.VectorSpace import Mathlib.Topology.MetricSpace.Perfect /-! diff --git a/Mathlib/Topology/Algebra/SeparationQuotient.lean b/Mathlib/Topology/Algebra/SeparationQuotient/Basic.lean similarity index 83% rename from Mathlib/Topology/Algebra/SeparationQuotient.lean rename to Mathlib/Topology/Algebra/SeparationQuotient/Basic.lean index 8a99f4dc7ef34..26a143140affc 100644 --- a/Mathlib/Topology/Algebra/SeparationQuotient.lean +++ b/Mathlib/Topology/Algebra/SeparationQuotient/Basic.lean @@ -3,10 +3,7 @@ Copyright (c) 2024 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ -import Mathlib.LinearAlgebra.Basis.VectorSpace -import Mathlib.Algebra.Module.Projective import Mathlib.Topology.Algebra.Module.Basic -import Mathlib.Topology.Maps.OpenQuotient /-! # Algebraic operations on `SeparationQuotient` @@ -22,6 +19,8 @@ Finally, we construct a section of the quotient map which is a continuous linear map `SeparationQuotient E →L[K] E`. -/ +assert_not_exists LinearIndependent + namespace SeparationQuotient section SMul @@ -397,79 +396,4 @@ theorem mk_algebraMap (r : R) : mk (algebraMap R A r) = algebraMap R (Separation end Algebra -section VectorSpace - -variable (K E : Type*) [DivisionRing K] [AddCommGroup E] [Module K E] - [TopologicalSpace E] [TopologicalAddGroup E] [ContinuousConstSMul K E] - -/-- There exists a continuous `K`-linear map from `SeparationQuotient E` to `E` -such that `mk (outCLM x) = x` for all `x`. - -Note that continuity of this map comes for free, because `mk` is a topology inducing map. --/ -theorem exists_out_continuousLinearMap : - ∃ f : SeparationQuotient E →L[K] E, mkCLM K E ∘L f = .id K (SeparationQuotient E) := by - rcases (mkCLM K E).toLinearMap.exists_rightInverse_of_surjective - (LinearMap.range_eq_top.mpr surjective_mk) with ⟨f, hf⟩ - replace hf : mk ∘ f = id := congr_arg DFunLike.coe hf - exact ⟨⟨f, inducing_mk.continuous_iff.2 (by continuity)⟩, DFunLike.ext' hf⟩ - -/-- A continuous `K`-linear map from `SeparationQuotient E` to `E` -such that `mk (outCLM x) = x` for all `x`. -/ -noncomputable def outCLM : SeparationQuotient E →L[K] E := - (exists_out_continuousLinearMap K E).choose - -@[simp] -theorem mkCLM_comp_outCLM : mkCLM K E ∘L outCLM K E = .id K (SeparationQuotient E) := - (exists_out_continuousLinearMap K E).choose_spec - -variable {E} in -@[simp] -theorem mk_outCLM (x : SeparationQuotient E) : mk (outCLM K E x) = x := - DFunLike.congr_fun (mkCLM_comp_outCLM K E) x - -@[simp] -theorem mk_comp_outCLM : mk ∘ outCLM K E = id := funext (mk_outCLM K) - -variable {K} in -theorem postcomp_mkCLM_surjective {L : Type*} [Semiring L] (σ : L →+* K) - (F : Type*) [AddCommMonoid F] [Module L F] [TopologicalSpace F] : - Function.Surjective ((mkCLM K E).comp : (F →SL[σ] E) → (F →SL[σ] SeparationQuotient E)) := by - intro f - use (outCLM K E).comp f - rw [← ContinuousLinearMap.comp_assoc, mkCLM_comp_outCLM, ContinuousLinearMap.id_comp] - -/-- The `SeparationQuotient.outCLM K E` map is a topological embedding. -/ -theorem outCLM_embedding : Embedding (outCLM K E) := - Function.LeftInverse.embedding (mk_outCLM K) continuous_mk (map_continuous _) - -theorem outCLM_injective : Function.Injective (outCLM K E) := - (outCLM_embedding K E).injective - -end VectorSpace - -section VectorSpaceUniform - -variable (K E : Type*) [DivisionRing K] [AddCommGroup E] [Module K E] - [UniformSpace E] [UniformAddGroup E] [ContinuousConstSMul K E] - -theorem outCLM_isUniformInducing : IsUniformInducing (outCLM K E) := by - rw [← isUniformInducing_mk.isUniformInducing_comp_iff, mk_comp_outCLM] - exact .id - -@[deprecated (since := "2024-10-05")] -alias outCLM_uniformInducing := outCLM_isUniformInducing - -theorem outCLM_isUniformEmbedding : IsUniformEmbedding (outCLM K E) where - inj := outCLM_injective K E - toIsUniformInducing := outCLM_isUniformInducing K E - -@[deprecated (since := "2024-10-01")] -alias outCLM_uniformEmbedding := outCLM_isUniformEmbedding - -theorem outCLM_uniformContinuous : UniformContinuous (outCLM K E) := - (outCLM_isUniformInducing K E).uniformContinuous - -end VectorSpaceUniform - end SeparationQuotient diff --git a/Mathlib/Topology/Algebra/SeparationQuotient/Section.lean b/Mathlib/Topology/Algebra/SeparationQuotient/Section.lean new file mode 100644 index 0000000000000..5e4941bd8a541 --- /dev/null +++ b/Mathlib/Topology/Algebra/SeparationQuotient/Section.lean @@ -0,0 +1,93 @@ +/- +Copyright (c) 2024 Yury Kudryashov. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yury Kudryashov +-/ +import Mathlib.Algebra.Module.Projective +import Mathlib.LinearAlgebra.Basis.VectorSpace +import Mathlib.Topology.Algebra.SeparationQuotient.Basic +import Mathlib.Topology.Maps.OpenQuotient + +/-! +# Algebraic operations on `SeparationQuotient` + +In this file we construct a section of the quotient map `E → SeparationQuotient E` as a continuous +linear map `SeparationQuotient E →L[K] E`. +-/ + +namespace SeparationQuotient +section VectorSpace + +variable (K E : Type*) [DivisionRing K] [AddCommGroup E] [Module K E] + [TopologicalSpace E] [TopologicalAddGroup E] [ContinuousConstSMul K E] + +/-- There exists a continuous `K`-linear map from `SeparationQuotient E` to `E` +such that `mk (outCLM x) = x` for all `x`. + +Note that continuity of this map comes for free, because `mk` is a topology inducing map. +-/ +theorem exists_out_continuousLinearMap : + ∃ f : SeparationQuotient E →L[K] E, mkCLM K E ∘L f = .id K (SeparationQuotient E) := by + rcases (mkCLM K E).toLinearMap.exists_rightInverse_of_surjective + (LinearMap.range_eq_top.mpr surjective_mk) with ⟨f, hf⟩ + replace hf : mk ∘ f = id := congr_arg DFunLike.coe hf + exact ⟨⟨f, inducing_mk.continuous_iff.2 (by continuity)⟩, DFunLike.ext' hf⟩ + +/-- A continuous `K`-linear map from `SeparationQuotient E` to `E` +such that `mk (outCLM x) = x` for all `x`. -/ +noncomputable def outCLM : SeparationQuotient E →L[K] E := + (exists_out_continuousLinearMap K E).choose + +@[simp] +theorem mkCLM_comp_outCLM : mkCLM K E ∘L outCLM K E = .id K (SeparationQuotient E) := + (exists_out_continuousLinearMap K E).choose_spec + +variable {E} in +@[simp] +theorem mk_outCLM (x : SeparationQuotient E) : mk (outCLM K E x) = x := + DFunLike.congr_fun (mkCLM_comp_outCLM K E) x + +@[simp] +theorem mk_comp_outCLM : mk ∘ outCLM K E = id := funext (mk_outCLM K) + +variable {K} in +theorem postcomp_mkCLM_surjective {L : Type*} [Semiring L] (σ : L →+* K) + (F : Type*) [AddCommMonoid F] [Module L F] [TopologicalSpace F] : + Function.Surjective ((mkCLM K E).comp : (F →SL[σ] E) → (F →SL[σ] SeparationQuotient E)) := by + intro f + use (outCLM K E).comp f + rw [← ContinuousLinearMap.comp_assoc, mkCLM_comp_outCLM, ContinuousLinearMap.id_comp] + +/-- The `SeparationQuotient.outCLM K E` map is a topological embedding. -/ +theorem outCLM_embedding : Embedding (outCLM K E) := + Function.LeftInverse.embedding (mk_outCLM K) continuous_mk (map_continuous _) + +theorem outCLM_injective : Function.Injective (outCLM K E) := + (outCLM_embedding K E).injective + +end VectorSpace + +section VectorSpaceUniform + +variable (K E : Type*) [DivisionRing K] [AddCommGroup E] [Module K E] + [UniformSpace E] [UniformAddGroup E] [ContinuousConstSMul K E] + +theorem outCLM_isUniformInducing : IsUniformInducing (outCLM K E) := by + rw [← isUniformInducing_mk.isUniformInducing_comp_iff, mk_comp_outCLM] + exact .id + +@[deprecated (since := "2024-10-05")] +alias outCLM_uniformInducing := outCLM_isUniformInducing + +theorem outCLM_isUniformEmbedding : IsUniformEmbedding (outCLM K E) where + inj := outCLM_injective K E + toIsUniformInducing := outCLM_isUniformInducing K E + +@[deprecated (since := "2024-10-01")] +alias outCLM_uniformEmbedding := outCLM_isUniformEmbedding + +theorem outCLM_uniformContinuous : UniformContinuous (outCLM K E) := + (outCLM_isUniformInducing K E).uniformContinuous + +end VectorSpaceUniform +end SeparationQuotient diff --git a/Mathlib/Topology/MetricSpace/Algebra.lean b/Mathlib/Topology/MetricSpace/Algebra.lean index c3e28fc646228..e8a16f85a6d7a 100644 --- a/Mathlib/Topology/MetricSpace/Algebra.lean +++ b/Mathlib/Topology/MetricSpace/Algebra.lean @@ -4,9 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Heather Macbeth -/ import Mathlib.Topology.Algebra.MulAction +import Mathlib.Topology.Algebra.SeparationQuotient.Basic import Mathlib.Topology.Algebra.UniformMulAction import Mathlib.Topology.MetricSpace.Lipschitz -import Mathlib.Topology.Algebra.SeparationQuotient /-! # Compatibility of algebraic operations with metric space structures diff --git a/Mathlib/Topology/PartitionOfUnity.lean b/Mathlib/Topology/PartitionOfUnity.lean index bee6a23c8a177..d5fb0185564ef 100644 --- a/Mathlib/Topology/PartitionOfUnity.lean +++ b/Mathlib/Topology/PartitionOfUnity.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import Mathlib.Algebra.BigOperators.Finprod +import Mathlib.LinearAlgebra.Basis.VectorSpace import Mathlib.Topology.ContinuousMap.Algebra import Mathlib.Topology.Compactness.Paracompact import Mathlib.Topology.ShrinkingLemma From c3f00b0c794c8e9d0f75cef11354d6efa9a8ceea Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Mar=C3=ADa=20In=C3=A9s=20de=20Frutos-Fern=C3=A1ndez?= Date: Mon, 21 Oct 2024 18:56:49 +0000 Subject: [PATCH 248/505] feat(Analysis/SpecificLimits/Basic): add lemmas about limits of fractions (#16768) Used in #15373. Co-authored-by: Floris van Doorn --- Mathlib/Analysis/SpecificLimits/Basic.lean | 47 ++++++++++++++++++++++ Mathlib/Data/Set/Pointwise/Interval.lean | 12 ++++++ Mathlib/Topology/Algebra/Order/Field.lean | 34 ++++++++++++++++ 3 files changed, 93 insertions(+) diff --git a/Mathlib/Analysis/SpecificLimits/Basic.lean b/Mathlib/Analysis/SpecificLimits/Basic.lean index c0a2df7d23e61..e06aef0a9c56a 100644 --- a/Mathlib/Analysis/SpecificLimits/Basic.lean +++ b/Mathlib/Analysis/SpecificLimits/Basic.lean @@ -114,6 +114,53 @@ theorem tendsto_natCast_div_add_atTop {𝕜 : Type*} [DivisionRing 𝕜] [Topolo intros simp_all only [comp_apply, map_inv₀, map_natCast] +/-- For any positive `m : ℕ`, `((n % m : ℕ) : ℝ) / (n : ℝ)` tends to `0` as `n` tends to `∞`. -/ +theorem tendsto_mod_div_atTop_nhds_zero_nat {m : ℕ} (hm : 0 < m) : + Tendsto (fun n : ℕ => ((n % m : ℕ) : ℝ) / (n : ℝ)) atTop (𝓝 0) := by + have h0 : ∀ᶠ n : ℕ in atTop, 0 ≤ (fun n : ℕ => ((n % m : ℕ) : ℝ)) n := by aesop + exact tendsto_bdd_div_atTop_nhds_zero h0 + (.of_forall (fun n ↦ cast_le.mpr (mod_lt n hm).le)) tendsto_natCast_atTop_atTop + +theorem Filter.EventuallyEq.div_mul_cancel {α G : Type*} [GroupWithZero G] {f g : α → G} + {l : Filter α} (hg : Tendsto g l (𝓟 {0}ᶜ)) : (fun x ↦ f x / g x * g x) =ᶠ[l] fun x ↦ f x := by + filter_upwards [hg.le_comap <| preimage_mem_comap (m := g) (mem_principal_self {0}ᶜ)] with x hx + aesop + +/-- If `g` tends to `∞`, then eventually for all `x` we have `(f x / g x) * g x = f x`. -/ +theorem Filter.EventuallyEq.div_mul_cancel_atTop {α K : Type*} [LinearOrderedSemifield K] + {f g : α → K} {l : Filter α} (hg : Tendsto g l atTop) : + (fun x ↦ f x / g x * g x) =ᶠ[l] fun x ↦ f x := + div_mul_cancel <| hg.mono_right <| le_principal_iff.mpr <| + mem_of_superset (Ioi_mem_atTop 0) <| by aesop + +/-- If when `x` tends to `∞`, `g` tends to `∞` and `f x / g x` tends to a positive + constant, then `f` tends to `∞`. -/ +theorem Tendsto.num {α K : Type*} [LinearOrderedField K] [TopologicalSpace K] [OrderTopology K] + {f g : α → K} {l : Filter α} (hg : Tendsto g l atTop) {a : K} (ha : 0 < a) + (hlim : Tendsto (fun x => f x / g x) l (𝓝 a)) : + Tendsto f l atTop := + Tendsto.congr' (EventuallyEq.div_mul_cancel_atTop hg) (Tendsto.mul_atTop ha hlim hg) + +/-- If when `x` tends to `∞`, `g` tends to `∞` and `f x / g x` tends to a positive + constant, then `f` tends to `∞`. -/ +theorem Tendsto.den {α K : Type*} [LinearOrderedField K] [TopologicalSpace K] [OrderTopology K] + [ContinuousInv K] {f g : α → K} {l : Filter α} (hf : Tendsto f l atTop) {a : K} (ha : 0 < a) + (hlim : Tendsto (fun x => f x / g x) l (𝓝 a)) : + Tendsto g l atTop := by + have hlim' : Tendsto (fun x => g x / f x) l (𝓝 a⁻¹) := by + simp_rw [← inv_div (f _)] + exact Filter.Tendsto.inv (f := fun x => f x / g x) hlim + apply Tendsto.congr' (EventuallyEq.div_mul_cancel_atTop hf) + (Tendsto.mul_atTop (inv_pos_of_pos ha) hlim' hf) + +/-- If when `x` tends to `∞`, `f x / g x` tends to a positive constant, then `f` tends to `∞` if + and only if `g` tends to `∞`. -/ +theorem Tendsto.num_atTop_iff_den_atTop {α K : Type*} [LinearOrderedField K] [TopologicalSpace K] + [OrderTopology K] [ContinuousInv K] {f g : α → K} {l : Filter α} {a : K} (ha : 0 < a) + (hlim : Tendsto (fun x => f x / g x) l (𝓝 a)) : + Tendsto f l atTop ↔ Tendsto g l atTop := + ⟨fun hf ↦ Tendsto.den hf ha hlim, fun hg ↦ Tendsto.num hg ha hlim⟩ + /-! ### Powers -/ diff --git a/Mathlib/Data/Set/Pointwise/Interval.lean b/Mathlib/Data/Set/Pointwise/Interval.lean index 96c008f454acd..88b3b2c811895 100644 --- a/Mathlib/Data/Set/Pointwise/Interval.lean +++ b/Mathlib/Data/Set/Pointwise/Interval.lean @@ -670,6 +670,18 @@ theorem preimage_div_const_uIcc (ha : a ≠ 0) (b c : α) : (fun x => x / a) ⁻¹' [[b, c]] = [[b * a, c * a]] := by simp only [div_eq_mul_inv, preimage_mul_const_uIcc (inv_ne_zero ha), inv_inv] +lemma preimage_const_mul_Ioi_or_Iio (hb : a ≠ 0) {U V : Set α} + (hU : U ∈ {s | ∃ a, s = Ioi a ∨ s = Iio a}) (hV : V = HMul.hMul a ⁻¹' U) : + V ∈ {s | ∃ a, s = Ioi a ∨ s = Iio a} := by + obtain ⟨aU, (haU | haU)⟩ := hU <;> + simp only [hV, haU, mem_setOf_eq] <;> + use a⁻¹ * aU <;> + rcases lt_or_gt_of_ne hb with (hb | hb) + · right; rw [Set.preimage_const_mul_Ioi_of_neg _ hb, div_eq_inv_mul] + · left; rw [Set.preimage_const_mul_Ioi _ hb, div_eq_inv_mul] + · left; rw [Set.preimage_const_mul_Iio_of_neg _ hb, div_eq_inv_mul] + · right; rw [Set.preimage_const_mul_Iio _ hb, div_eq_inv_mul] + @[simp] theorem image_mul_const_uIcc (a b c : α) : (· * a) '' [[b, c]] = [[b * a, c * a]] := if ha : a = 0 then by simp [ha] diff --git a/Mathlib/Topology/Algebra/Order/Field.lean b/Mathlib/Topology/Algebra/Order/Field.lean index b80ccdfcef320..8e6a51ab5faed 100644 --- a/Mathlib/Topology/Algebra/Order/Field.lean +++ b/Mathlib/Topology/Algebra/Order/Field.lean @@ -141,6 +141,40 @@ theorem Filter.Tendsto.inv_tendsto_atTop (h : Tendsto f l atTop) : Tendsto f⁻ theorem Filter.Tendsto.inv_tendsto_zero (h : Tendsto f l (𝓝[>] 0)) : Tendsto f⁻¹ l atTop := tendsto_inv_zero_atTop.comp h +/-- If `g` tends to zero and there exists a constant `C : 𝕜` such that eventually `|f x| ≤ C`, + then the product `f * g` tends to zero. -/ +theorem bdd_le_mul_tendsto_zero' {f g : α → 𝕜} (C : 𝕜) (hf : ∀ᶠ x in l, |f x| ≤ C) + (hg : Tendsto g l (𝓝 0)) : Tendsto (fun x ↦ f x * g x) l (𝓝 0) := by + rw [tendsto_zero_iff_abs_tendsto_zero] + have hC : Tendsto (fun x ↦ |C * g x|) l (𝓝 0) := by + convert (hg.const_mul C).abs + simp_rw [mul_zero, abs_zero] + apply tendsto_of_tendsto_of_tendsto_of_le_of_le' tendsto_const_nhds hC + · filter_upwards [hf] with x _ using abs_nonneg _ + · filter_upwards [hf] with x hx + simp only [comp_apply, abs_mul] + exact mul_le_mul_of_nonneg_right (hx.trans (le_abs_self C)) (abs_nonneg _) + +/-- If `g` tends to zero and there exist constants `b B : 𝕜` such that eventually `b ≤ f x| ≤ B`, + then the product `f * g` tends to zero. -/ +theorem bdd_le_mul_tendsto_zero {f g : α → 𝕜} {b B : 𝕜} (hb : ∀ᶠ x in l, b ≤ f x) + (hB : ∀ᶠ x in l, f x ≤ B) (hg : Tendsto g l (𝓝 0)) : + Tendsto (fun x ↦ f x * g x) l (𝓝 0) := by + set C := max |b| |B| + have hbC : -C ≤ b := neg_le.mpr (le_max_of_le_left (neg_le_abs b)) + have hBC : B ≤ C := le_max_of_le_right (le_abs_self B) + apply bdd_le_mul_tendsto_zero' C _ hg + filter_upwards [hb, hB] + exact fun x hbx hBx ↦ abs_le.mpr ⟨hbC.trans hbx, hBx.trans hBC⟩ + +/-- If `g` tends to `atTop` and there exist constants `b B : 𝕜` such that eventually + `b ≤ f x| ≤ B`, then the quotient `f / g` tends to zero. -/ +theorem tendsto_bdd_div_atTop_nhds_zero {f g : α → 𝕜} {b B : 𝕜} + (hb : ∀ᶠ x in l, b ≤ f x) (hB : ∀ᶠ x in l, f x ≤ B) (hg : Tendsto g l atTop) : + Tendsto (fun x => f x / g x) l (𝓝 0) := by + simp only [div_eq_mul_inv] + exact bdd_le_mul_tendsto_zero hb hB hg.inv_tendsto_atTop + /-- The function `x^(-n)` tends to `0` at `+∞` for any positive natural `n`. A version for positive real powers exists as `tendsto_rpow_neg_atTop`. -/ theorem tendsto_pow_neg_atTop {n : ℕ} (hn : n ≠ 0) : From 13d0934bbc04883bc4edb521cbbb24920db26fb7 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 21 Oct 2024 21:50:36 +0000 Subject: [PATCH 249/505] =?UTF-8?q?feat:=20`=E2=80=96aba=E2=81=BB=C2=B9b?= =?UTF-8?q?=E2=81=BB=C2=B9=20-=201=E2=80=96=20=E2=89=A4=20C=20=E2=80=96a?= =?UTF-8?q?=20-=201=E2=80=96=20=E2=80=96b=20-=201=E2=80=96`=20(#18027)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit for `a`, `b` invertible. This will used to show the Breuillard-Green-Tao theorem in the case of matrices, which will then be bootstrapped to the full Breuillard-Green-Tao theorem. From [GrowthInGroups](https://github.com/YaelDillies/LeanCamCombi/tree/master/LeanCamCombi/GrowthInGroups) (LeanCamCombi) --- Mathlib/Analysis/Normed/Field/Basic.lean | 27 +++++++++++++++++++++++- 1 file changed, 26 insertions(+), 1 deletion(-) diff --git a/Mathlib/Analysis/Normed/Field/Basic.lean b/Mathlib/Analysis/Normed/Field/Basic.lean index 346a70a6a9cc2..4368eb6b8b5c3 100644 --- a/Mathlib/Analysis/Normed/Field/Basic.lean +++ b/Mathlib/Analysis/Normed/Field/Basic.lean @@ -454,6 +454,23 @@ lemma nnnorm_sub_mul_le (ha : ‖a‖₊ ≤ 1) : ‖c - a * b‖₊ ≤ ‖c - chord length is a metric on the unit complex numbers. -/ lemma nnnorm_sub_mul_le' (hb : ‖b‖₊ ≤ 1) : ‖c - a * b‖₊ ≤ ‖1 - a‖₊ + ‖c - b‖₊ := norm_sub_mul_le' hb +lemma norm_commutator_units_sub_one_le (a b : αˣ) : + ‖(a * b * a⁻¹ * b⁻¹).val - 1‖ ≤ 2 * ‖a⁻¹.val‖ * ‖b⁻¹.val‖ * ‖a.val - 1‖ * ‖b.val - 1‖ := + calc + ‖(a * b * a⁻¹ * b⁻¹).val - 1‖ = ‖(a * b - b * a) * a⁻¹.val * b⁻¹.val‖ := by simp [sub_mul, *] + _ ≤ ‖(a * b - b * a : α)‖ * ‖a⁻¹.val‖ * ‖b⁻¹.val‖ := norm_mul₃_le + _ = ‖(a - 1 : α) * (b - 1) - (b - 1) * (a - 1)‖ * ‖a⁻¹.val‖ * ‖b⁻¹.val‖ := by + simp_rw [sub_one_mul, mul_sub_one]; abel_nf + _ ≤ (‖(a - 1 : α) * (b - 1)‖ + ‖(b - 1 : α) * (a - 1)‖) * ‖a⁻¹.val‖ * ‖b⁻¹.val‖ := by + gcongr; exact norm_sub_le .. + _ ≤ (‖a.val - 1‖ * ‖b.val - 1‖ + ‖b.val - 1‖ * ‖a.val - 1‖) * ‖a⁻¹.val‖ * ‖b⁻¹.val‖ := by + gcongr <;> exact norm_mul_le .. + _ = 2 * ‖a⁻¹.val‖ * ‖b⁻¹.val‖ * ‖a.val - 1‖ * ‖b.val - 1‖ := by ring + +lemma nnnorm_commutator_units_sub_one_le (a b : αˣ) : + ‖(a * b * a⁻¹ * b⁻¹).val - 1‖₊ ≤ 2 * ‖a⁻¹.val‖₊ * ‖b⁻¹.val‖₊ * ‖a.val - 1‖₊ * ‖b.val - 1‖₊ := by + simpa using norm_commutator_units_sub_one_le a b + /-- A homomorphism `f` between semi_normed_rings is bounded if there exists a positive constant `C` such that for all `x` in `α`, `norm (f x) ≤ C * norm x`. -/ def RingHom.IsBounded {α : Type*} [SeminormedRing α] {β : Type*} [SeminormedRing β] @@ -613,7 +630,7 @@ end NormedCommRing section NormedDivisionRing -variable [NormedDivisionRing α] {a : α} +variable [NormedDivisionRing α] {a b : α} @[simp] theorem norm_mul (a b : α) : ‖a * b‖ = ‖a‖ * ‖b‖ := @@ -695,6 +712,14 @@ theorem nndist_inv_inv₀ {z w : α} (hz : z ≠ 0) (hw : w ≠ 0) : nndist z⁻¹ w⁻¹ = nndist z w / (‖z‖₊ * ‖w‖₊) := NNReal.eq <| dist_inv_inv₀ hz hw +lemma norm_commutator_sub_one_le (ha : a ≠ 0) (hb : b ≠ 0) : + ‖a * b * a⁻¹ * b⁻¹ - 1‖ ≤ 2 * ‖a‖⁻¹ * ‖b‖⁻¹ * ‖a - 1‖ * ‖b - 1‖ := by + simpa using norm_commutator_units_sub_one_le (.mk0 a ha) (.mk0 b hb) + +lemma nnnorm_commutator_sub_one_le (ha : a ≠ 0) (hb : b ≠ 0) : + ‖a * b * a⁻¹ * b⁻¹ - 1‖₊ ≤ 2 * ‖a‖₊⁻¹ * ‖b‖₊⁻¹ * ‖a - 1‖₊ * ‖b - 1‖₊ := by + simpa using nnnorm_commutator_units_sub_one_le (.mk0 a ha) (.mk0 b hb) + namespace NormedDivisionRing section Discrete From ef1ba70ef8afacfe951dc751154db2675a6530e3 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 21 Oct 2024 22:54:16 +0000 Subject: [PATCH 250/505] refactor: merge `ExpChar` and `CharP` API (#18023) `ExpChar` is just as basic as `CharP` but was defined in a later file importing a lot of linear algebra. This situation resulted resulted in a lot of import bloat further down the line. I personally noticed it because `LinearAlgebra.Span` was imported in `Data.Nat.Totient`! This PR merges the `ExpChar` and `CharP` into `Algebra.CharP.Defs` and `Algebra.CharP.Basic`. This lets us prove some `CharP` API using the corresponding `ExpChar` one, and ensures that the two APIs match up in lemmas. Some lemmas were moved in or out of the `CharP` namespace for consistency (but I am to keep the statu quo for now if you ask) The only thing left in `Algebra.CharP.ExpChar`, is the Frobenius endomorphism bundled as a linear map. I am not sure where to put it. It can be moved later. --- Mathlib/Algebra/CharP/Algebra.lean | 29 ++ Mathlib/Algebra/CharP/Basic.lean | 325 +++++++++++++--- Mathlib/Algebra/CharP/Defs.lean | 145 ++++++++ Mathlib/Algebra/CharP/ExpChar.lean | 352 +----------------- Mathlib/Algebra/CharP/IntermediateField.lean | 4 +- Mathlib/Algebra/CharP/LocalRing.lean | 1 - Mathlib/Algebra/CharP/Reduced.lean | 2 +- Mathlib/Algebra/CharP/Two.lean | 4 +- Mathlib/Algebra/Polynomial/Expand.lean | 1 + Mathlib/Data/Nat/Choose/Lucas.lean | 3 +- Mathlib/Data/Nat/Totient.lean | 3 + Mathlib/FieldTheory/Finite/GaloisField.lean | 2 +- Mathlib/FieldTheory/PerfectClosure.lean | 1 + Mathlib/FieldTheory/PurelyInseparable.lean | 3 +- Mathlib/FieldTheory/SeparableDegree.lean | 4 +- Mathlib/NumberTheory/MulChar/Basic.lean | 2 +- Mathlib/RingTheory/FiniteType.lean | 1 + Mathlib/RingTheory/Perfection.lean | 1 + Mathlib/RingTheory/Polynomial/Basic.lean | 3 +- .../Polynomial/SeparableDegree.lean | 1 - 20 files changed, 472 insertions(+), 415 deletions(-) diff --git a/Mathlib/Algebra/CharP/Algebra.lean b/Mathlib/Algebra/CharP/Algebra.lean index 05641583b0b03..7d73460eb9f55 100644 --- a/Mathlib/Algebra/CharP/Algebra.lean +++ b/Mathlib/Algebra/CharP/Algebra.lean @@ -67,6 +67,35 @@ theorem RingHom.charP_iff {R A : Type*} [NonAssocSemiring R] [NonAssocSemiring A (H : Function.Injective f) (p : ℕ) : CharP R p ↔ CharP A p := ⟨fun _ ↦ charP_of_injective_ringHom H p, fun _ ↦ f.charP H p⟩ +/-- If a ring homomorphism `R →+* A` is injective then `A` has the same exponential characteristic +as `R`. -/ +lemma expChar_of_injective_ringHom {R A : Type*} + [Semiring R] [Semiring A] {f : R →+* A} (h : Function.Injective f) + (q : ℕ) [hR : ExpChar R q] : ExpChar A q := by + rcases hR with _ | hprime + · haveI := charZero_of_injective_ringHom h; exact .zero + haveI := charP_of_injective_ringHom h q; exact .prime hprime + +/-- If `R →+* A` is injective, and `A` is of exponential characteristic `p`, then `R` is also of +exponential characteristic `p`. Similar to `RingHom.charZero`. -/ +lemma RingHom.expChar {R A : Type*} [Semiring R] [Semiring A] (f : R →+* A) + (H : Function.Injective f) (p : ℕ) [ExpChar A p] : ExpChar R p := by + cases ‹ExpChar A p› with + | zero => haveI := f.charZero; exact .zero + | prime hp => haveI := f.charP H p; exact .prime hp + +/-- If `R →+* A` is injective, then `R` is of exponential characteristic `p` if and only if `A` is +also of exponential characteristic `p`. Similar to `RingHom.charZero_iff`. -/ +lemma RingHom.expChar_iff {R A : Type*} [Semiring R] [Semiring A] (f : R →+* A) + (H : Function.Injective f) (p : ℕ) : ExpChar R p ↔ ExpChar A p := + ⟨fun _ ↦ expChar_of_injective_ringHom H p, fun _ ↦ f.expChar H p⟩ + +/-- If the algebra map `R →+* A` is injective then `A` has the same exponential characteristic +as `R`. -/ +lemma expChar_of_injective_algebraMap {R A : Type*} [CommSemiring R] [Semiring A] [Algebra R A] + (h : Function.Injective (algebraMap R A)) (q : ℕ) [ExpChar R q] : ExpChar A q := + expChar_of_injective_ringHom h q + /-! As an application, a `ℚ`-algebra has characteristic zero. -/ diff --git a/Mathlib/Algebra/CharP/Basic.lean b/Mathlib/Algebra/CharP/Basic.lean index 6440a87f6d324..271aa60a6d144 100644 --- a/Mathlib/Algebra/CharP/Basic.lean +++ b/Mathlib/Algebra/CharP/Basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau, Joey van Langen, Casper Putz -/ import Mathlib.Algebra.CharP.Defs +import Mathlib.Algebra.GroupPower.IterateHom import Mathlib.Data.Nat.Multiplicity import Mathlib.Data.Nat.Choose.Sum @@ -11,13 +12,13 @@ import Mathlib.Data.Nat.Choose.Sum # Characteristic of semirings -/ +assert_not_exists Algebra +assert_not_exists LinearMap assert_not_exists orderOf -universe u v - open Finset -variable {R : Type*} +variable {R S : Type*} namespace Commute @@ -79,75 +80,172 @@ theorem exists_add_pow_prime_eq : end CommSemiring -variable (R) (x y : R) (n : ℕ) +section Semiring +variable [Semiring R] {x y : R} (p n : ℕ) -theorem add_pow_char_of_commute [Semiring R] {p : ℕ} [hp : Fact p.Prime] [CharP R p] - (h : Commute x y) : (x + y) ^ p = x ^ p + y ^ p := by - let ⟨r, hr⟩ := h.exists_add_pow_prime_eq hp.out - simp [hr] +section ExpChar +variable [hR : ExpChar R p] -theorem add_pow_char_pow_of_commute [Semiring R] {p : ℕ} [hp : Fact p.Prime] [CharP R p] - (h : Commute x y) : (x + y) ^ p ^ n = x ^ p ^ n + y ^ p ^ n := by - let ⟨r, hr⟩ := h.exists_add_pow_prime_pow_eq hp.out n - simp [hr] +lemma add_pow_expChar_of_commute (h : Commute x y) : (x + y) ^ p = x ^ p + y ^ p := by + obtain _ | hprime := hR + · simp only [pow_one] + · let ⟨r, hr⟩ := h.exists_add_pow_prime_eq hprime + simp [hr] -theorem sub_pow_char_of_commute [Ring R] {p : ℕ} [Fact p.Prime] [CharP R p] (h : Commute x y) : - (x - y) ^ p = x ^ p - y ^ p := by - rw [eq_sub_iff_add_eq, ← add_pow_char_of_commute _ _ _ (Commute.sub_left h rfl)] - simp +lemma add_pow_expChar_pow_of_commute (h : Commute x y) : + (x + y) ^ p ^ n = x ^ p ^ n + y ^ p ^ n := by + obtain _ | hprime := hR + · simp only [one_pow, pow_one] + · let ⟨r, hr⟩ := h.exists_add_pow_prime_pow_eq hprime n + simp [hr] -theorem sub_pow_char_pow_of_commute [Ring R] {p : ℕ} [Fact p.Prime] [CharP R p] (h : Commute x y) : - (x - y) ^ p ^ n = x ^ p ^ n - y ^ p ^ n := by - induction n with - | zero => simp - | succ n n_ih => - rw [pow_succ, pow_mul, pow_mul, pow_mul, n_ih] - apply sub_pow_char_of_commute; apply Commute.pow_pow h - -theorem add_pow_char [CommSemiring R] {p : ℕ} [Fact p.Prime] [CharP R p] : - (x + y) ^ p = x ^ p + y ^ p := - add_pow_char_of_commute _ _ _ (Commute.all _ _) - -theorem add_pow_char_pow [CommSemiring R] {p : ℕ} [Fact p.Prime] [CharP R p] : - (x + y) ^ p ^ n = x ^ p ^ n + y ^ p ^ n := - add_pow_char_pow_of_commute _ _ _ _ (Commute.all _ _) - -theorem add_pow_eq_add_pow_mod_mul_pow_add_pow_div - [CommSemiring R] {p : ℕ} [Fact p.Prime] [CharP R p] (x y : R) : +lemma add_pow_eq_mul_pow_add_pow_div_expChar_of_commute (h : Commute x y) : (x + y) ^ n = (x + y) ^ (n % p) * (x ^ p + y ^ p) ^ (n / p) := by - rw [← add_pow_char, ← pow_mul, ← pow_add, Nat.mod_add_div] + rw [← add_pow_expChar_of_commute _ h, ← pow_mul, ← pow_add, Nat.mod_add_div] + +end ExpChar + +section CharP +variable [hp : Fact p.Prime] [CharP R p] + +lemma add_pow_char_of_commute (h : Commute x y) : (x + y) ^ p = x ^ p + y ^ p := + add_pow_expChar_of_commute _ h + +lemma add_pow_char_pow_of_commute (h : Commute x y) : (x + y) ^ p ^ n = x ^ p ^ n + y ^ p ^ n := + add_pow_expChar_pow_of_commute _ _ h + +lemma add_pow_eq_mul_pow_add_pow_div_char_of_commute (h : Commute x y) : + (x + y) ^ n = (x + y) ^ (n % p) * (x ^ p + y ^ p) ^ (n / p) := + add_pow_eq_mul_pow_add_pow_div_expChar_of_commute _ _ h + +end CharP +end Semiring + +section CommSemiring +variable [CommSemiring R] (x y : R) (p n : ℕ) + +section ExpChar +variable [hR : ExpChar R p] + +lemma add_pow_expChar : (x + y) ^ p = x ^ p + y ^ p := add_pow_expChar_of_commute _ <| .all .. + +lemma add_pow_expChar_pow : (x + y) ^ p ^ n = x ^ p ^ n + y ^ p ^ n := + add_pow_expChar_pow_of_commute _ _ <| .all .. + +lemma add_pow_eq_mul_pow_add_pow_div_expChar : + (x + y) ^ n = (x + y) ^ (n % p) * (x ^ p + y ^ p) ^ (n / p) := + add_pow_eq_mul_pow_add_pow_div_expChar_of_commute _ _ <| .all .. + +end ExpChar + +section CharP +variable [hp : Fact p.Prime] [CharP R p] + +lemma add_pow_char : (x + y) ^ p = x ^ p + y ^ p := add_pow_expChar .. +lemma add_pow_char_pow : (x + y) ^ p ^ n = x ^ p ^ n + y ^ p ^ n := add_pow_expChar_pow .. + +lemma add_pow_eq_mul_pow_add_pow_div_char : + (x + y) ^ n = (x + y) ^ (n % p) * (x ^ p + y ^ p) ^ (n / p) := + add_pow_eq_mul_pow_add_pow_div_expChar .. + +@[deprecated (since := "2024-10-21")] +alias add_pow_eq_add_pow_mod_mul_pow_add_pow_div := add_pow_eq_mul_pow_add_pow_div_char + +end CharP +end CommSemiring -theorem sub_pow_char [CommRing R] {p : ℕ} [Fact p.Prime] [CharP R p] : - (x - y) ^ p = x ^ p - y ^ p := - sub_pow_char_of_commute _ _ _ (Commute.all _ _) +section Ring +variable [Ring R] {x y : R} (p n : ℕ) + +section ExpChar +variable [hR : ExpChar R p] +include hR + +lemma sub_pow_expChar_of_commute (h : Commute x y) : (x - y) ^ p = x ^ p - y ^ p := by + simp [eq_sub_iff_add_eq, ← add_pow_expChar_of_commute _ (h.sub_left rfl)] -theorem sub_pow_char_pow [CommRing R] {p : ℕ} [Fact p.Prime] [CharP R p] : - (x - y) ^ p ^ n = x ^ p ^ n - y ^ p ^ n := - sub_pow_char_pow_of_commute _ _ _ _ (Commute.all _ _) +lemma sub_pow_expChar_pow_of_commute (h : Commute x y) : + (x - y) ^ p ^ n = x ^ p ^ n - y ^ p ^ n := by + simp [eq_sub_iff_add_eq, ← add_pow_expChar_pow_of_commute _ _ (h.sub_left rfl)] -theorem sub_pow_eq_sub_pow_mod_mul_pow_sub_pow_div [CommRing R] {p : ℕ} [Fact p.Prime] [CharP R p] : +lemma sub_pow_eq_mul_pow_sub_pow_div_expChar_of_commute (h : Commute x y) : (x - y) ^ n = (x - y) ^ (n % p) * (x ^ p - y ^ p) ^ (n / p) := by - rw [← sub_pow_char, ← pow_mul, ← pow_add, Nat.mod_add_div] + rw [← sub_pow_expChar_of_commute _ h, ← pow_mul, ← pow_add, Nat.mod_add_div] + +variable (R) -theorem CharP.neg_one_pow_char [Ring R] (p : ℕ) [Fact p.Prime] [CharP R p] : - (-1 : R) ^ p = -1 := by +lemma neg_one_pow_expChar : (-1 : R) ^ p = -1 := by rw [eq_neg_iff_add_eq_zero] nth_rw 2 [← one_pow p] - rw [← add_pow_char_of_commute R _ _ (Commute.one_right _), neg_add_cancel, - zero_pow (Fact.out (p := Nat.Prime p)).ne_zero] + rw [← add_pow_expChar_of_commute _ (Commute.one_right _), neg_add_cancel, + zero_pow (expChar_ne_zero R p)] -theorem CharP.neg_one_pow_char_pow [Ring R] (p : ℕ) [CharP R p] [Fact p.Prime] : - (-1 : R) ^ p ^ n = -1 := by +lemma neg_one_pow_expChar_pow : (-1 : R) ^ p ^ n = -1 := by rw [eq_neg_iff_add_eq_zero] nth_rw 2 [← one_pow (p ^ n)] - rw [← add_pow_char_pow_of_commute R _ _ _ (Commute.one_right _), neg_add_cancel, - zero_pow (pow_ne_zero _ (Fact.out (p := Nat.Prime p)).ne_zero)] + rw [← add_pow_expChar_pow_of_commute _ _ (Commute.one_right _), neg_add_cancel, + zero_pow (pow_ne_zero _ <| expChar_ne_zero R p)] + +end ExpChar + +section CharP +variable [hp : Fact p.Prime] [CharP R p] + +lemma sub_pow_char_of_commute (h : Commute x y) : (x - y) ^ p = x ^ p - y ^ p := + sub_pow_expChar_of_commute _ h + +lemma sub_pow_char_pow_of_commute (h : Commute x y) : (x - y) ^ p ^ n = x ^ p ^ n - y ^ p ^ n := + sub_pow_expChar_pow_of_commute _ _ h + +variable (R) + +lemma neg_one_pow_char : (-1 : R) ^ p = -1 := neg_one_pow_expChar .. + +lemma neg_one_pow_char_pow : (-1 : R) ^ p ^ n = -1 := neg_one_pow_expChar_pow .. + +lemma sub_pow_eq_mul_pow_sub_pow_div_char_of_commute (h : Commute x y) : + (x - y) ^ n = (x - y) ^ (n % p) * (x ^ p - y ^ p) ^ (n / p) := + sub_pow_eq_mul_pow_sub_pow_div_expChar_of_commute _ _ h + +end CharP +end Ring + +section CommRing +variable [CommRing R] (x y : R) (n : ℕ) {p : ℕ} + +section ExpChar +variable [hR : ExpChar R p] + +lemma sub_pow_expChar : (x - y) ^ p = x ^ p - y ^ p := sub_pow_expChar_of_commute _ <| .all .. + +lemma sub_pow_expChar_pow : (x - y) ^ p ^ n = x ^ p ^ n - y ^ p ^ n := + sub_pow_expChar_pow_of_commute _ _ <| .all .. + +lemma sub_pow_eq_mul_pow_sub_pow_div_expChar : + (x - y) ^ n = (x - y) ^ (n % p) * (x ^ p - y ^ p) ^ (n / p) := + sub_pow_eq_mul_pow_sub_pow_div_expChar_of_commute _ _ <| .all .. + +end ExpChar + +section CharP +variable [hp : Fact p.Prime] [CharP R p] + +lemma sub_pow_char : (x - y) ^ p = x ^ p - y ^ p := sub_pow_expChar .. +lemma sub_pow_char_pow : (x - y) ^ p ^ n = x ^ p ^ n - y ^ p ^ n := sub_pow_expChar_pow .. + +lemma sub_pow_eq_mul_pow_sub_pow_div_char : + (x - y) ^ n = (x - y) ^ (n % p) * (x ^ p - y ^ p) ^ (n / p) := + sub_pow_eq_mul_pow_sub_pow_div_expChar .. + +end CharP +end CommRing + namespace CharP section -variable [NonAssocRing R] +variable (R) [NonAssocRing R] /-- The characteristic of a finite ring cannot be zero. -/ theorem char_ne_zero_of_finite (p : ℕ) [CharP R p] [Finite R] : p ≠ 0 := by @@ -163,10 +261,131 @@ end section Ring -variable [Ring R] [NoZeroDivisors R] [Nontrivial R] [Finite R] +variable (R) [Ring R] [NoZeroDivisors R] [Nontrivial R] [Finite R] theorem char_is_prime (p : ℕ) [CharP R p] : p.Prime := Or.resolve_right (char_is_prime_or_zero R p) (char_ne_zero_of_finite R p) end Ring end CharP + +/-! ### The Frobenius automorphism -/ + +section frobenius +section CommSemiring +variable [CommSemiring R] {S : Type*} [CommSemiring S] (f : R →* S) (g : R →+* S) (p m n : ℕ) + [ExpChar R p] [ExpChar S p] (x y : R) + +open ExpChar + +variable (R) in +/-- The frobenius map `x ↦ x ^ p`. -/ +def frobenius : R →+* R where + __ := powMonoidHom p + map_zero' := zero_pow (expChar_pos R p).ne' + map_add' _ _ := add_pow_expChar .. + +variable (R) in +/-- The iterated frobenius map `x ↦ x ^ p ^ n`. -/ +def iterateFrobenius : R →+* R where + __ := powMonoidHom (p ^ n) + map_zero' := zero_pow (expChar_pow_pos R p n).ne' + map_add' _ _ := add_pow_expChar_pow .. + + +lemma frobenius_def : frobenius R p x = x ^ p := rfl + +lemma iterateFrobenius_def : iterateFrobenius R p n x = x ^ p ^ n := rfl + +lemma iterate_frobenius : (frobenius R p)^[n] x = x ^ p ^ n := congr_fun (pow_iterate p n) x + +variable (R) + +lemma coe_iterateFrobenius : iterateFrobenius R p n = (frobenius R p)^[n] := + (pow_iterate p n).symm + +lemma iterateFrobenius_one_apply : iterateFrobenius R p 1 x = x ^ p := by + rw [iterateFrobenius_def, pow_one] + +@[simp] +lemma iterateFrobenius_one : iterateFrobenius R p 1 = frobenius R p := + RingHom.ext (iterateFrobenius_one_apply R p) + +lemma iterateFrobenius_zero_apply : iterateFrobenius R p 0 x = x := by + rw [iterateFrobenius_def, pow_zero, pow_one] + +@[simp] +lemma iterateFrobenius_zero : iterateFrobenius R p 0 = RingHom.id R := + RingHom.ext (iterateFrobenius_zero_apply R p) + +lemma iterateFrobenius_add_apply : + iterateFrobenius R p (m + n) x = iterateFrobenius R p m (iterateFrobenius R p n x) := by + simp_rw [iterateFrobenius_def, add_comm m n, pow_add, pow_mul] + +lemma iterateFrobenius_add : + iterateFrobenius R p (m + n) = (iterateFrobenius R p m).comp (iterateFrobenius R p n) := + RingHom.ext (iterateFrobenius_add_apply R p m n) + +lemma iterateFrobenius_mul_apply : + iterateFrobenius R p (m * n) x = (iterateFrobenius R p m)^[n] x := by + simp_rw [coe_iterateFrobenius, Function.iterate_mul] + +lemma coe_iterateFrobenius_mul : iterateFrobenius R p (m * n) = (iterateFrobenius R p m)^[n] := + funext (iterateFrobenius_mul_apply R p m n) + +variable {R} + +lemma frobenius_mul : frobenius R p (x * y) = frobenius R p x * frobenius R p y := + map_mul (frobenius R p) x y + +lemma frobenius_one : frobenius R p 1 = 1 := one_pow _ + +lemma MonoidHom.map_frobenius : f (frobenius R p x) = frobenius S p (f x) := map_pow f x p +lemma RingHom.map_frobenius : g (frobenius R p x) = frobenius S p (g x) := map_pow g x p + +lemma MonoidHom.map_iterate_frobenius (n : ℕ) : + f ((frobenius R p)^[n] x) = (frobenius S p)^[n] (f x) := + Function.Semiconj.iterate_right (f.map_frobenius p) n x + +lemma RingHom.map_iterate_frobenius (n : ℕ) : + g ((frobenius R p)^[n] x) = (frobenius S p)^[n] (g x) := + g.toMonoidHom.map_iterate_frobenius p x n + +lemma MonoidHom.iterate_map_frobenius (f : R →* R) (p : ℕ) [ExpChar R p] (n : ℕ) : + f^[n] (frobenius R p x) = frobenius R p (f^[n] x) := + iterate_map_pow f _ _ _ + +lemma RingHom.iterate_map_frobenius (f : R →+* R) (p : ℕ) [ExpChar R p] (n : ℕ) : + f^[n] (frobenius R p x) = frobenius R p (f^[n] x) := iterate_map_pow f _ _ _ + +lemma list_sum_pow_char (l : List R) : l.sum ^ p = (l.map (· ^ p : R → R)).sum := + map_list_sum (frobenius R p) _ + +lemma multiset_sum_pow_char (s : Multiset R) : s.sum ^ p = (s.map (· ^ p : R → R)).sum := + map_multiset_sum (frobenius R p) _ + +lemma sum_pow_char {ι : Type*} (s : Finset ι) (f : ι → R) : (∑ i ∈ s, f i) ^ p = ∑ i ∈ s, f i ^ p := + map_sum (frobenius R p) _ _ + +variable (n : ℕ) + +lemma list_sum_pow_char_pow (l : List R) : l.sum ^ p ^ n = (l.map (· ^ p ^ n : R → R)).sum := + map_list_sum (iterateFrobenius R p n) _ + +lemma multiset_sum_pow_char_pow (s : Multiset R) : + s.sum ^ p ^ n = (s.map (· ^ p ^ n : R → R)).sum := map_multiset_sum (iterateFrobenius R p n) _ + +lemma sum_pow_char_pow {ι : Type*} (s : Finset ι) (f : ι → R) : + (∑ i ∈ s, f i) ^ p ^ n = ∑ i ∈ s, f i ^ p ^ n := map_sum (iterateFrobenius R p n) _ _ + +end CommSemiring + +section CommRing +variable [CommRing R] (p : ℕ) [ExpChar R p] (x y : R) + +lemma frobenius_neg : frobenius R p (-x) = -frobenius R p x := map_neg .. + +lemma frobenius_sub : frobenius R p (x - y) = frobenius R p x - frobenius R p y := map_sub .. + +end CommRing +end frobenius diff --git a/Mathlib/Algebra/CharP/Defs.lean b/Mathlib/Algebra/CharP/Defs.lean index 3f42840a6923c..9509007537ca2 100644 --- a/Mathlib/Algebra/CharP/Defs.lean +++ b/Mathlib/Algebra/CharP/Defs.lean @@ -272,6 +272,10 @@ lemma char_is_prime_or_zero (p : ℕ) [hc : CharP R p] : Nat.Prime p ∨ p = 0 : | 1, hc => absurd (Eq.refl (1 : ℕ)) (@char_ne_one R _ _ (1 : ℕ) hc) | m + 2, hc => Or.inl (@char_is_prime_of_two_le R _ _ (m + 2) hc (Nat.le_add_left 2 m)) +/-- The characteristic is prime if it is non-zero. -/ +lemma char_prime_of_ne_zero {p : ℕ} [CharP R p] (hp : p ≠ 0) : p.Prime := + (CharP.char_is_prime_or_zero R p).resolve_right hp + lemma exists' (R : Type*) [NonAssocRing R] [NoZeroDivisors R] [Nontrivial R] : CharZero R ∨ ∃ p : ℕ, Fact p.Prime ∧ CharP R p := by obtain ⟨p, hchar⟩ := CharP.exists R @@ -436,3 +440,144 @@ namespace Fin instance charP (n : ℕ) [NeZero n] : CharP (Fin n) n where cast_eq_zero_iff' _ := natCast_eq_zero end Fin + +/-! +### Exponential characteristic + +This section defines the exponential characteristic, which is defined to be 1 for a ring with +characteristic 0 and the same as the ordinary characteristic, if the ordinary characteristic is +prime. This concept is useful to simplify some theorem statements. +This file establishes a few basic results relating it to the (ordinary characteristic). +The definition is stated for a semiring, but the actual results are for nontrivial rings +(as far as exponential characteristic one is concerned), respectively a ring without zero-divisors +(for prime characteristic). +-/ + +section AddMonoidWithOne +variable [AddMonoidWithOne R] + +/-- The definition of the exponential characteristic of a semiring. -/ +class inductive ExpChar : ℕ → Prop + | zero [CharZero R] : ExpChar 1 + | prime {q : ℕ} (hprime : q.Prime) [hchar : CharP R q] : ExpChar q + +instance expChar_prime (p) [CharP R p] [Fact p.Prime] : ExpChar R p := ExpChar.prime Fact.out +instance expChar_one [CharZero R] : ExpChar R 1 := ExpChar.zero + +lemma expChar_ne_zero (p : ℕ) [hR : ExpChar R p] : p ≠ 0 := by + cases hR + · exact one_ne_zero + · exact ‹p.Prime›.ne_zero + +instance (S : Type*) [Semiring S] (p) [ExpChar R p] [ExpChar S p] : ExpChar (R × S) p := by + obtain hp | ⟨hp⟩ := ‹ExpChar R p› + · have := Prod.charZero_of_left R S; exact .zero + obtain _ | _ := ‹ExpChar S p› + · exact (Nat.not_prime_one hp).elim + · have := Prod.charP R S p; exact .prime hp + +variable {R} in +/-- The exponential characteristic is unique. -/ +lemma ExpChar.eq {p q : ℕ} (hp : ExpChar R p) (hq : ExpChar R q) : p = q := by + rcases hp with ⟨hp⟩ | ⟨hp'⟩ + · rcases hq with hq | hq' + exacts [rfl, False.elim (Nat.not_prime_zero (CharP.eq R ‹_› (CharP.ofCharZero R) ▸ hq'))] + · rcases hq with hq | hq' + exacts [False.elim (Nat.not_prime_zero (CharP.eq R ‹_› (CharP.ofCharZero R) ▸ hp')), + CharP.eq R ‹_› ‹_›] + +lemma ExpChar.congr {p : ℕ} (q : ℕ) [hq : ExpChar R q] (h : q = p) : ExpChar R p := h ▸ hq + +/-- The exponential characteristic is one if the characteristic is zero. -/ +lemma expChar_one_of_char_zero (q : ℕ) [hp : CharP R 0] [hq : ExpChar R q] : q = 1 := by + rcases hq with q | hq_prime + · rfl + · exact False.elim <| hq_prime.ne_zero <| ‹CharP R q›.eq R hp + +/-- The characteristic equals the exponential characteristic iff the former is prime. -/ +lemma char_eq_expChar_iff (p q : ℕ) [hp : CharP R p] [hq : ExpChar R q] : p = q ↔ p.Prime := by + rcases hq with q | hq_prime + · rw [(CharP.eq R hp inferInstance : p = 0)] + decide + · exact ⟨fun hpq => hpq.symm ▸ hq_prime, fun _ => CharP.eq R hp ‹CharP R q›⟩ + +/-- The exponential characteristic is a prime number or one. +See also `CharP.char_is_prime_or_zero`. -/ +lemma expChar_is_prime_or_one (q : ℕ) [hq : ExpChar R q] : Nat.Prime q ∨ q = 1 := by + cases hq with + | zero => exact .inr rfl + | prime hp => exact .inl hp + +/-- The exponential characteristic is positive. -/ +lemma expChar_pos (q : ℕ) [ExpChar R q] : 0 < q := by + rcases expChar_is_prime_or_one R q with h | rfl + exacts [Nat.Prime.pos h, Nat.one_pos] + +/-- Any power of the exponential characteristic is positive. -/ +lemma expChar_pow_pos (q : ℕ) [ExpChar R q] (n : ℕ) : 0 < q ^ n := + Nat.pos_pow_of_pos n (expChar_pos R q) + +end AddMonoidWithOne + +section NonAssocSemiring +variable [NonAssocSemiring R] + +/-- Noncomputable function that outputs the unique exponential characteristic of a semiring. -/ +noncomputable def ringExpChar : ℕ := max (ringChar R) 1 + +lemma ringExpChar.eq (q : ℕ) [h : ExpChar R q] : ringExpChar R = q := by + rcases h with _ | h + · haveI := CharP.ofCharZero R + rw [ringExpChar, ringChar.eq R 0]; rfl + rw [ringExpChar, ringChar.eq R q] + exact Nat.max_eq_left h.one_lt.le + +@[simp] lemma ringExpChar.eq_one [CharZero R] : ringExpChar R = 1 := by + rw [ringExpChar, ringChar.eq_zero, max_eq_right (Nat.zero_le _)] + +section Nontrivial +variable [Nontrivial R] + +/-- The exponential characteristic is one if the characteristic is zero. -/ +lemma char_zero_of_expChar_one (p : ℕ) [hp : CharP R p] [hq : ExpChar R 1] : p = 0 := by + cases hq + · exact CharP.eq R hp inferInstance + · exact False.elim (CharP.char_ne_one R 1 rfl) + +-- This could be an instance, but there are no `ExpChar R 1` instances in mathlib. +/-- The characteristic is zero if the exponential characteristic is one. -/ +lemma charZero_of_expChar_one' [hq : ExpChar R 1] : CharZero R := by + cases hq + · assumption + · exact False.elim (CharP.char_ne_one R 1 rfl) + +/-- The exponential characteristic is one iff the characteristic is zero. -/ +lemma expChar_one_iff_char_zero (p q : ℕ) [CharP R p] [ExpChar R q] : q = 1 ↔ p = 0 := by + constructor + · rintro rfl + exact char_zero_of_expChar_one R p + · rintro rfl + exact expChar_one_of_char_zero R q + +end Nontrivial +end NonAssocSemiring + +lemma ExpChar.exists [Ring R] [IsDomain R] : ∃ q, ExpChar R q := by + obtain _ | ⟨p, ⟨hp⟩, _⟩ := CharP.exists' R + exacts [⟨1, .zero⟩, ⟨p, .prime hp⟩] + +lemma ExpChar.exists_unique [Ring R] [IsDomain R] : ∃! q, ExpChar R q := + let ⟨q, H⟩ := ExpChar.exists R + ⟨q, H, fun _ H2 ↦ ExpChar.eq H2 H⟩ + +instance ringExpChar.expChar [Ring R] [IsDomain R] : ExpChar R (ringExpChar R) := by + obtain ⟨q, _⟩ := ExpChar.exists R + rwa [ringExpChar.eq R q] + +variable {R} in +lemma ringExpChar.of_eq [Ring R] [IsDomain R] {q : ℕ} (h : ringExpChar R = q) : ExpChar R q := + h ▸ ringExpChar.expChar R + +variable {R} in +lemma ringExpChar.eq_iff [Ring R] [IsDomain R] {q : ℕ} : ringExpChar R = q ↔ ExpChar R q := + ⟨ringExpChar.of_eq, fun _ ↦ ringExpChar.eq R q⟩ diff --git a/Mathlib/Algebra/CharP/ExpChar.lean b/Mathlib/Algebra/CharP/ExpChar.lean index 24ac40fe6a5e2..ebc8479fc69c9 100644 --- a/Mathlib/Algebra/CharP/ExpChar.lean +++ b/Mathlib/Algebra/CharP/ExpChar.lean @@ -3,9 +3,8 @@ Copyright (c) 2021 Jakob Scholbach. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jakob Scholbach -/ +import Mathlib.Algebra.Algebra.Defs import Mathlib.Algebra.CharP.Basic -import Mathlib.Algebra.CharP.Algebra -import Mathlib.Data.Nat.Prime.Defs /-! # Exponential characteristic @@ -28,235 +27,12 @@ The definition is stated for a semiring, but the actual results are for nontrivi exponential characteristic, characteristic -/ +open ExpChar universe u variable (R : Type u) -section Semiring - -variable [Semiring R] - -/-- The definition of the exponential characteristic of a semiring. -/ -class inductive ExpChar (R : Type u) [Semiring R] : ℕ → Prop - | zero [CharZero R] : ExpChar R 1 - | prime {q : ℕ} (hprime : q.Prime) [hchar : CharP R q] : ExpChar R q - -instance expChar_prime (p) [CharP R p] [Fact p.Prime] : ExpChar R p := ExpChar.prime Fact.out -instance expChar_zero [CharZero R] : ExpChar R 1 := ExpChar.zero - -instance (S : Type*) [Semiring S] (p) [ExpChar R p] [ExpChar S p] : ExpChar (R × S) p := by - obtain hp | ⟨hp⟩ := ‹ExpChar R p› - · have := Prod.charZero_of_left R S; exact .zero - obtain _ | _ := ‹ExpChar S p› - · exact (Nat.not_prime_one hp).elim - · have := Prod.charP R S p; exact .prime hp - -variable {R} in -/-- The exponential characteristic is unique. -/ -theorem ExpChar.eq {p q : ℕ} (hp : ExpChar R p) (hq : ExpChar R q) : p = q := by - rcases hp with ⟨hp⟩ | ⟨hp'⟩ - · rcases hq with hq | hq' - exacts [rfl, False.elim (Nat.not_prime_zero (CharP.eq R ‹_› (CharP.ofCharZero R) ▸ hq'))] - · rcases hq with hq | hq' - exacts [False.elim (Nat.not_prime_zero (CharP.eq R ‹_› (CharP.ofCharZero R) ▸ hp')), - CharP.eq R ‹_› ‹_›] - -theorem ExpChar.congr {p : ℕ} (q : ℕ) [hq : ExpChar R q] (h : q = p) : ExpChar R p := h ▸ hq - -/-- Noncomputable function that outputs the unique exponential characteristic of a semiring. -/ -noncomputable def ringExpChar (R : Type*) [NonAssocSemiring R] : ℕ := max (ringChar R) 1 - -theorem ringExpChar.eq (q : ℕ) [h : ExpChar R q] : ringExpChar R = q := by - rcases h with _ | h - · haveI := CharP.ofCharZero R - rw [ringExpChar, ringChar.eq R 0]; rfl - rw [ringExpChar, ringChar.eq R q] - exact Nat.max_eq_left h.one_lt.le - -@[simp] -theorem ringExpChar.eq_one (R : Type*) [NonAssocSemiring R] [CharZero R] : ringExpChar R = 1 := by - rw [ringExpChar, ringChar.eq_zero, max_eq_right zero_le_one] - -/-- The exponential characteristic is one if the characteristic is zero. -/ -theorem expChar_one_of_char_zero (q : ℕ) [hp : CharP R 0] [hq : ExpChar R q] : q = 1 := by - rcases hq with q | hq_prime - · rfl - · exact False.elim <| hq_prime.ne_zero <| ‹CharP R q›.eq R hp - -/-- The characteristic equals the exponential characteristic iff the former is prime. -/ -theorem char_eq_expChar_iff (p q : ℕ) [hp : CharP R p] [hq : ExpChar R q] : p = q ↔ p.Prime := by - rcases hq with q | hq_prime - · rw [(CharP.eq R hp inferInstance : p = 0)] - decide - · exact ⟨fun hpq => hpq.symm ▸ hq_prime, fun _ => CharP.eq R hp ‹CharP R q›⟩ - -/-- The exponential characteristic is a prime number or one. -See also `CharP.char_is_prime_or_zero`. -/ -theorem expChar_is_prime_or_one (q : ℕ) [hq : ExpChar R q] : Nat.Prime q ∨ q = 1 := by - cases hq with - | zero => exact .inr rfl - | prime hp => exact .inl hp - -/-- The exponential characteristic is positive. -/ -theorem expChar_pos (q : ℕ) [ExpChar R q] : 0 < q := by - rcases expChar_is_prime_or_one R q with h | rfl - exacts [Nat.Prime.pos h, Nat.one_pos] - -/-- Any power of the exponential characteristic is positive. -/ -theorem expChar_pow_pos (q : ℕ) [ExpChar R q] (n : ℕ) : 0 < q ^ n := - Nat.pos_pow_of_pos n (expChar_pos R q) - -section Nontrivial - -variable [Nontrivial R] - -/-- The exponential characteristic is one if the characteristic is zero. -/ -theorem char_zero_of_expChar_one (p : ℕ) [hp : CharP R p] [hq : ExpChar R 1] : p = 0 := by - cases hq - · exact CharP.eq R hp inferInstance - · exact False.elim (CharP.char_ne_one R 1 rfl) - --- This could be an instance, but there are no `ExpChar R 1` instances in mathlib. -/-- The characteristic is zero if the exponential characteristic is one. -/ -theorem charZero_of_expChar_one' [hq : ExpChar R 1] : CharZero R := by - cases hq - · assumption - · exact False.elim (CharP.char_ne_one R 1 rfl) - -/-- The exponential characteristic is one iff the characteristic is zero. -/ -theorem expChar_one_iff_char_zero (p q : ℕ) [CharP R p] [ExpChar R q] : q = 1 ↔ p = 0 := by - constructor - · rintro rfl - exact char_zero_of_expChar_one R p - · rintro rfl - exact expChar_one_of_char_zero R q - -section NoZeroDivisors - -variable [NoZeroDivisors R] - -/-- A helper lemma: the characteristic is prime if it is non-zero. -/ -theorem char_prime_of_ne_zero {p : ℕ} [hp : CharP R p] (p_ne_zero : p ≠ 0) : Nat.Prime p := by - rcases CharP.char_is_prime_or_zero R p with h | h - · exact h - · contradiction - -end NoZeroDivisors - -end Nontrivial - -end Semiring - -theorem ExpChar.exists [Ring R] [IsDomain R] : ∃ q, ExpChar R q := by - obtain _ | ⟨p, ⟨hp⟩, _⟩ := CharP.exists' R - exacts [⟨1, .zero⟩, ⟨p, .prime hp⟩] - -theorem ExpChar.exists_unique [Ring R] [IsDomain R] : ∃! q, ExpChar R q := - let ⟨q, H⟩ := ExpChar.exists R - ⟨q, H, fun _ H2 ↦ ExpChar.eq H2 H⟩ - -instance ringExpChar.expChar [Ring R] [IsDomain R] : ExpChar R (ringExpChar R) := by - obtain ⟨q, _⟩ := ExpChar.exists R - rwa [ringExpChar.eq R q] - -variable {R} in -theorem ringExpChar.of_eq [Ring R] [IsDomain R] {q : ℕ} (h : ringExpChar R = q) : ExpChar R q := - h ▸ ringExpChar.expChar R - -variable {R} in -theorem ringExpChar.eq_iff [Ring R] [IsDomain R] {q : ℕ} : ringExpChar R = q ↔ ExpChar R q := - ⟨ringExpChar.of_eq, fun _ ↦ ringExpChar.eq R q⟩ - -/-- If a ring homomorphism `R →+* A` is injective then `A` has the same exponential characteristic -as `R`. -/ -theorem expChar_of_injective_ringHom {R A : Type*} - [Semiring R] [Semiring A] {f : R →+* A} (h : Function.Injective f) - (q : ℕ) [hR : ExpChar R q] : ExpChar A q := by - rcases hR with _ | hprime - · haveI := charZero_of_injective_ringHom h; exact .zero - haveI := charP_of_injective_ringHom h q; exact .prime hprime - -/-- If `R →+* A` is injective, and `A` is of exponential characteristic `p`, then `R` is also of -exponential characteristic `p`. Similar to `RingHom.charZero`. -/ -theorem RingHom.expChar {R A : Type*} [Semiring R] [Semiring A] (f : R →+* A) - (H : Function.Injective f) (p : ℕ) [ExpChar A p] : ExpChar R p := by - cases ‹ExpChar A p› with - | zero => haveI := f.charZero; exact .zero - | prime hp => haveI := f.charP H p; exact .prime hp - -/-- If `R →+* A` is injective, then `R` is of exponential characteristic `p` if and only if `A` is -also of exponential characteristic `p`. Similar to `RingHom.charZero_iff`. -/ -theorem RingHom.expChar_iff {R A : Type*} [Semiring R] [Semiring A] (f : R →+* A) - (H : Function.Injective f) (p : ℕ) : ExpChar R p ↔ ExpChar A p := - ⟨fun _ ↦ expChar_of_injective_ringHom H p, fun _ ↦ f.expChar H p⟩ - -/-- If the algebra map `R →+* A` is injective then `A` has the same exponential characteristic -as `R`. -/ -theorem expChar_of_injective_algebraMap {R A : Type*} - [CommSemiring R] [Semiring A] [Algebra R A] (h : Function.Injective (algebraMap R A)) - (q : ℕ) [ExpChar R q] : ExpChar A q := expChar_of_injective_ringHom h q - -theorem add_pow_expChar_of_commute [Semiring R] {q : ℕ} [hR : ExpChar R q] - (x y : R) (h : Commute x y) : (x + y) ^ q = x ^ q + y ^ q := by - rcases hR with _ | hprime - · simp only [pow_one] - haveI := Fact.mk hprime; exact add_pow_char_of_commute R x y h - -theorem add_pow_expChar_pow_of_commute [Semiring R] {q : ℕ} [hR : ExpChar R q] - {n : ℕ} (x y : R) (h : Commute x y) : (x + y) ^ q ^ n = x ^ q ^ n + y ^ q ^ n := by - rcases hR with _ | hprime - · simp only [one_pow, pow_one] - haveI := Fact.mk hprime; exact add_pow_char_pow_of_commute R x y n h - -theorem sub_pow_expChar_of_commute [Ring R] {q : ℕ} [hR : ExpChar R q] - (x y : R) (h : Commute x y) : (x - y) ^ q = x ^ q - y ^ q := by - rcases hR with _ | hprime - · simp only [pow_one] - haveI := Fact.mk hprime; exact sub_pow_char_of_commute R x y h - -theorem sub_pow_expChar_pow_of_commute [Ring R] {q : ℕ} [hR : ExpChar R q] - {n : ℕ} (x y : R) (h : Commute x y) : (x - y) ^ q ^ n = x ^ q ^ n - y ^ q ^ n := by - rcases hR with _ | hprime - · simp only [one_pow, pow_one] - haveI := Fact.mk hprime; exact sub_pow_char_pow_of_commute R x y n h - -theorem add_pow_expChar [CommSemiring R] {q : ℕ} [hR : ExpChar R q] - (x y : R) : (x + y) ^ q = x ^ q + y ^ q := by - rcases hR with _ | hprime - · simp only [pow_one] - haveI := Fact.mk hprime; exact add_pow_char R x y - -theorem add_pow_expChar_pow [CommSemiring R] {q : ℕ} [hR : ExpChar R q] - {n : ℕ} (x y : R) : (x + y) ^ q ^ n = x ^ q ^ n + y ^ q ^ n := by - rcases hR with _ | hprime - · simp only [one_pow, pow_one] - haveI := Fact.mk hprime; exact add_pow_char_pow R x y n - -theorem sub_pow_expChar [CommRing R] {q : ℕ} [hR : ExpChar R q] - (x y : R) : (x - y) ^ q = x ^ q - y ^ q := by - rcases hR with _ | hprime - · simp only [pow_one] - haveI := Fact.mk hprime; exact sub_pow_char R x y - -theorem sub_pow_expChar_pow [CommRing R] {q : ℕ} [hR : ExpChar R q] - {n : ℕ} (x y : R) : (x - y) ^ q ^ n = x ^ q ^ n - y ^ q ^ n := by - rcases hR with _ | hprime - · simp only [one_pow, pow_one] - haveI := Fact.mk hprime; exact sub_pow_char_pow R x y n - -theorem ExpChar.neg_one_pow_expChar [Ring R] (q : ℕ) [hR : ExpChar R q] : - (-1 : R) ^ q = -1 := by - rcases hR with _ | hprime - · simp only [pow_one] - haveI := Fact.mk hprime; exact CharP.neg_one_pow_char R q - -theorem ExpChar.neg_one_pow_expChar_pow [Ring R] (q n : ℕ) [hR : ExpChar R q] : - (-1 : R) ^ q ^ n = -1 := by - rcases hR with _ | hprime - · simp only [one_pow, pow_one] - haveI := Fact.mk hprime; exact CharP.neg_one_pow_char_pow R n q - section frobenius section CommSemiring @@ -264,91 +40,7 @@ section CommSemiring variable [CommSemiring R] {S : Type*} [CommSemiring S] (f : R →* S) (g : R →+* S) (p m n : ℕ) [ExpChar R p] [ExpChar S p] (x y : R) -/-- The frobenius map that sends x to x^p -/ -def frobenius : R →+* R where - __ := powMonoidHom p - map_zero' := zero_pow (expChar_pos R p).ne' - map_add' := add_pow_expChar R - -/-- The iterated frobenius map sending x to x^p^n -/ -def iterateFrobenius : R →+* R where - __ := powMonoidHom (p ^ n) - map_zero' := zero_pow (expChar_pow_pos R p n).ne' - map_add' := add_pow_expChar_pow R - -variable {R} - -theorem frobenius_def : frobenius R p x = x ^ p := rfl - -theorem iterateFrobenius_def : iterateFrobenius R p n x = x ^ p ^ n := rfl - -theorem iterate_frobenius : (frobenius R p)^[n] x = x ^ p ^ n := congr_fun (pow_iterate p n) x - -variable (R) - -theorem coe_iterateFrobenius : iterateFrobenius R p n = (frobenius R p)^[n] := - (pow_iterate p n).symm - -theorem iterateFrobenius_one_apply : iterateFrobenius R p 1 x = x ^ p := by - rw [iterateFrobenius_def, pow_one] - -@[simp] -theorem iterateFrobenius_one : iterateFrobenius R p 1 = frobenius R p := - RingHom.ext (iterateFrobenius_one_apply R p) - -theorem iterateFrobenius_zero_apply : iterateFrobenius R p 0 x = x := by - rw [iterateFrobenius_def, pow_zero, pow_one] - -@[simp] -theorem iterateFrobenius_zero : iterateFrobenius R p 0 = RingHom.id R := - RingHom.ext (iterateFrobenius_zero_apply R p) - -theorem iterateFrobenius_add_apply : - iterateFrobenius R p (m + n) x = iterateFrobenius R p m (iterateFrobenius R p n x) := by - simp_rw [iterateFrobenius_def, add_comm m n, pow_add, pow_mul] - -theorem iterateFrobenius_add : - iterateFrobenius R p (m + n) = (iterateFrobenius R p m).comp (iterateFrobenius R p n) := - RingHom.ext (iterateFrobenius_add_apply R p m n) - -theorem iterateFrobenius_mul_apply : - iterateFrobenius R p (m * n) x = (iterateFrobenius R p m)^[n] x := by - simp_rw [coe_iterateFrobenius, Function.iterate_mul] - -theorem coe_iterateFrobenius_mul : iterateFrobenius R p (m * n) = (iterateFrobenius R p m)^[n] := - funext (iterateFrobenius_mul_apply R p m n) - -variable {R} - -theorem frobenius_mul : frobenius R p (x * y) = frobenius R p x * frobenius R p y := - map_mul (frobenius R p) x y - -theorem frobenius_one : frobenius R p 1 = 1 := - one_pow _ - -theorem MonoidHom.map_frobenius : f (frobenius R p x) = frobenius S p (f x) := - map_pow f x p - -theorem RingHom.map_frobenius : g (frobenius R p x) = frobenius S p (g x) := - map_pow g x p - -theorem MonoidHom.map_iterate_frobenius (n : ℕ) : - f ((frobenius R p)^[n] x) = (frobenius S p)^[n] (f x) := - Function.Semiconj.iterate_right (f.map_frobenius p) n x - -theorem RingHom.map_iterate_frobenius (n : ℕ) : - g ((frobenius R p)^[n] x) = (frobenius S p)^[n] (g x) := - g.toMonoidHom.map_iterate_frobenius p x n - -theorem MonoidHom.iterate_map_frobenius (f : R →* R) (p : ℕ) [ExpChar R p] (n : ℕ) : - f^[n] (frobenius R p x) = frobenius R p (f^[n] x) := - iterate_map_pow f _ _ _ - -theorem RingHom.iterate_map_frobenius (f : R →+* R) (p : ℕ) [ExpChar R p] (n : ℕ) : - f^[n] (frobenius R p x) = frobenius R p (f^[n] x) := - iterate_map_pow f _ _ _ - -variable (R S) +variable (S) /-- The frobenius map of an algebra as a frobenius-semilinear map. -/ nonrec def LinearMap.frobenius [Algebra R S] : S →ₛₗ[frobenius R p] S where @@ -379,43 +71,5 @@ theorem frobenius_natCast (n : ℕ) : frobenius R p n = n := @[deprecated (since := "2024-04-17")] alias frobenius_nat_cast := frobenius_natCast -variable {R} - -theorem list_sum_pow_char (l : List R) : l.sum ^ p = (l.map (· ^ p : R → R)).sum := - map_list_sum (frobenius R p) _ - -theorem multiset_sum_pow_char (s : Multiset R) : s.sum ^ p = (s.map (· ^ p : R → R)).sum := - map_multiset_sum (frobenius R p) _ - -theorem sum_pow_char {ι : Type*} (s : Finset ι) (f : ι → R) : - (∑ i ∈ s, f i) ^ p = ∑ i ∈ s, f i ^ p := - map_sum (frobenius R p) _ _ - -variable (n : ℕ) - -theorem list_sum_pow_char_pow (l : List R) : l.sum ^ p ^ n = (l.map (· ^ p ^ n : R → R)).sum := - map_list_sum (iterateFrobenius R p n) _ - -theorem multiset_sum_pow_char_pow (s : Multiset R) : - s.sum ^ p ^ n = (s.map (· ^ p ^ n : R → R)).sum := - map_multiset_sum (iterateFrobenius R p n) _ - -theorem sum_pow_char_pow {ι : Type*} (s : Finset ι) (f : ι → R) : - (∑ i ∈ s, f i) ^ p ^ n = ∑ i ∈ s, f i ^ p ^ n := - map_sum (iterateFrobenius R p n) _ _ - end CommSemiring - -section CommRing - -variable [CommRing R] (p : ℕ) [ExpChar R p] (x y : R) - -theorem frobenius_neg : frobenius R p (-x) = -frobenius R p x := - map_neg .. - -theorem frobenius_sub : frobenius R p (x - y) = frobenius R p x - frobenius R p y := - map_sub .. - -end CommRing - end frobenius diff --git a/Mathlib/Algebra/CharP/IntermediateField.lean b/Mathlib/Algebra/CharP/IntermediateField.lean index a577f00bd38d8..611dc38c84f62 100644 --- a/Mathlib/Algebra/CharP/IntermediateField.lean +++ b/Mathlib/Algebra/CharP/IntermediateField.lean @@ -3,9 +3,9 @@ Copyright (c) 2024 Jz Pan. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jz Pan -/ -import Mathlib.Algebra.CharP.ExpChar -import Mathlib.FieldTheory.IntermediateField.Basic +import Mathlib.Algebra.CharP.Algebra import Mathlib.Algebra.EuclideanDomain.Field +import Mathlib.FieldTheory.IntermediateField.Basic /-! diff --git a/Mathlib/Algebra/CharP/LocalRing.lean b/Mathlib/Algebra/CharP/LocalRing.lean index 645df7bf14817..224e3821a2e71 100644 --- a/Mathlib/Algebra/CharP/LocalRing.lean +++ b/Mathlib/Algebra/CharP/LocalRing.lean @@ -3,7 +3,6 @@ Copyright (c) 2022 Jon Eugster. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jon Eugster -/ -import Mathlib.Algebra.CharP.Basic import Mathlib.Algebra.IsPrimePow import Mathlib.Data.Nat.Factorization.Basic import Mathlib.RingTheory.LocalRing.ResidueField.Defs diff --git a/Mathlib/Algebra/CharP/Reduced.lean b/Mathlib/Algebra/CharP/Reduced.lean index 045af0fb1b707..7328d585042a2 100644 --- a/Mathlib/Algebra/CharP/Reduced.lean +++ b/Mathlib/Algebra/CharP/Reduced.lean @@ -3,7 +3,7 @@ Copyright (c) 2018 Kenny Lau. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau, Joey van Langen, Casper Putz -/ -import Mathlib.Algebra.CharP.ExpChar +import Mathlib.Algebra.CharP.Basic import Mathlib.RingTheory.Nilpotent.Defs /-! diff --git a/Mathlib/Algebra/CharP/Two.lean b/Mathlib/Algebra/CharP/Two.lean index 0f91ff1fa1e5b..6a4a8f2d12bda 100644 --- a/Mathlib/Algebra/CharP/Two.lean +++ b/Mathlib/Algebra/CharP/Two.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Eric Wieser. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Eric Wieser -/ -import Mathlib.Algebra.CharP.ExpChar +import Mathlib.Algebra.CharP.Basic import Mathlib.GroupTheory.OrderOfElement /-! @@ -15,6 +15,8 @@ The lemmas in this file with a `_sq` suffix are just special cases of the `_pow_ elsewhere, with a shorter name for ease of discovery, and no need for a `[Fact (Prime 2)]` argument. -/ +assert_not_exists Algebra +assert_not_exists LinearMap variable {R ι : Type*} diff --git a/Mathlib/Algebra/Polynomial/Expand.lean b/Mathlib/Algebra/Polynomial/Expand.lean index 33ed390954f87..25e32f8c28e20 100644 --- a/Mathlib/Algebra/Polynomial/Expand.lean +++ b/Mathlib/Algebra/Polynomial/Expand.lean @@ -3,6 +3,7 @@ Copyright (c) 2020 Kenny Lau. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau -/ +import Mathlib.Algebra.CharP.Basic import Mathlib.RingTheory.Polynomial.Basic /-! diff --git a/Mathlib/Data/Nat/Choose/Lucas.lean b/Mathlib/Data/Nat/Choose/Lucas.lean index 0df7d593456f8..0ada97a2cd2dd 100644 --- a/Mathlib/Data/Nat/Choose/Lucas.lean +++ b/Mathlib/Data/Nat/Choose/Lucas.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Gareth Ma. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Gareth Ma -/ +import Mathlib.Algebra.CharP.Basic import Mathlib.Data.ZMod.Basic import Mathlib.RingTheory.Polynomial.Basic @@ -33,7 +34,7 @@ modulo `p`. Also see `choose_modEq_choose_mod_mul_choose_div_nat` for the versio theorem choose_modEq_choose_mod_mul_choose_div : choose n k ≡ choose (n % p) (k % p) * choose (n / p) (k / p) [ZMOD p] := by have decompose : ((X : (ZMod p)[X]) + 1) ^ n = (X + 1) ^ (n % p) * (X ^ p + 1) ^ (n / p) := by - simpa using add_pow_eq_add_pow_mod_mul_pow_add_pow_div _ n (X : (ZMod p)[X]) 1 + simpa using add_pow_eq_mul_pow_add_pow_div_char (X : (ZMod p)[X]) 1 p _ simp only [← ZMod.intCast_eq_intCast_iff, Int.cast_mul, Int.cast_ofNat, ← coeff_X_add_one_pow _ n k, ← eq_intCast (Int.castRingHom (ZMod p)), ← coeff_map, Polynomial.map_pow, Polynomial.map_add, Polynomial.map_one, map_X, decompose] diff --git a/Mathlib/Data/Nat/Totient.lean b/Mathlib/Data/Nat/Totient.lean index ec08e85c219d9..59878b2a1d4a9 100644 --- a/Mathlib/Data/Nat/Totient.lean +++ b/Mathlib/Data/Nat/Totient.lean @@ -20,6 +20,9 @@ We prove the divisor sum formula, namely that `n` equals `φ` summed over the di `totient_prime_pow`. -/ +assert_not_exists Algebra +assert_not_exists LinearMap + open Finset namespace Nat diff --git a/Mathlib/FieldTheory/Finite/GaloisField.lean b/Mathlib/FieldTheory/Finite/GaloisField.lean index adbca6dc2715e..e3868fc90d591 100644 --- a/Mathlib/FieldTheory/Finite/GaloisField.lean +++ b/Mathlib/FieldTheory/Finite/GaloisField.lean @@ -131,7 +131,7 @@ theorem finrank {n} (h : n ≠ 0) : Module.finrank (ZMod p) (GaloisField p n) = rw [hx, hy] · intro x _ hx simp only [g_poly, sub_eq_zero, aeval_X_pow, aeval_X, map_sub, sub_neg_eq_add] at * - rw [neg_pow, hx, CharP.neg_one_pow_char_pow] + rw [neg_pow, hx, neg_one_pow_char_pow] simp · simp only [g_poly, aeval_X_pow, aeval_X, map_sub, mul_pow, sub_eq_zero] intro x y _ _ hx hy diff --git a/Mathlib/FieldTheory/PerfectClosure.lean b/Mathlib/FieldTheory/PerfectClosure.lean index 39f00a2095466..700ee9dc3791a 100644 --- a/Mathlib/FieldTheory/PerfectClosure.lean +++ b/Mathlib/FieldTheory/PerfectClosure.lean @@ -3,6 +3,7 @@ Copyright (c) 2018 Kenny Lau. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau, Yury Kudryashov -/ +import Mathlib.Algebra.CharP.ExpChar import Mathlib.FieldTheory.Perfect /-! diff --git a/Mathlib/FieldTheory/PurelyInseparable.lean b/Mathlib/FieldTheory/PurelyInseparable.lean index bf5f64619ceef..6099462a72603 100644 --- a/Mathlib/FieldTheory/PurelyInseparable.lean +++ b/Mathlib/FieldTheory/PurelyInseparable.lean @@ -3,8 +3,9 @@ Copyright (c) 2024 Jz Pan. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jz Pan -/ -import Mathlib.FieldTheory.SeparableClosure +import Mathlib.Algebra.CharP.ExpChar import Mathlib.Algebra.CharP.IntermediateField +import Mathlib.FieldTheory.SeparableClosure /-! diff --git a/Mathlib/FieldTheory/SeparableDegree.lean b/Mathlib/FieldTheory/SeparableDegree.lean index bc340b3686eef..ec3dbd6aa4fb7 100644 --- a/Mathlib/FieldTheory/SeparableDegree.lean +++ b/Mathlib/FieldTheory/SeparableDegree.lean @@ -639,10 +639,10 @@ theorem natSepDegree_eq_one_iff_eq_X_sub_C_pow : (minpoly F x).natSepDegree = 1 rw [minpoly.aeval, map_sub, map_pow, aeval_X, aeval_C, sub_eq_zero, eq_comm] at hx use n rw [h, Polynomial.map_sub, Polynomial.map_pow, map_X, map_C, hx, map_pow, - ← sub_pow_expChar_pow_of_commute E[X] X (C x) (commute_X _)] + ← sub_pow_expChar_pow_of_commute _ _ (commute_X _)] apply_fun constantCoeff at h simp_rw [map_pow, map_sub, constantCoeff_apply, coeff_map, coeff_X_zero, coeff_C_zero] at h - rw [zero_sub, neg_pow, ExpChar.neg_one_pow_expChar_pow] at h + rw [zero_sub, neg_pow, neg_one_pow_expChar_pow] at h exact ⟨n, -(minpoly F x).coeff 0, by rw [map_neg, h, neg_mul, one_mul, neg_neg]⟩ end minpoly diff --git a/Mathlib/NumberTheory/MulChar/Basic.lean b/Mathlib/NumberTheory/MulChar/Basic.lean index e4ed0932f9876..827d48e679919 100644 --- a/Mathlib/NumberTheory/MulChar/Basic.lean +++ b/Mathlib/NumberTheory/MulChar/Basic.lean @@ -502,7 +502,7 @@ theorem IsQuadratic.pow_char {χ : MulChar R R'} (hχ : χ.IsQuadratic) (p : ℕ rcases hχ x with (hx | hx | hx) <;> rw [hx] · rw [zero_pow (@Fact.out p.Prime).ne_zero] · rw [one_pow] - · exact CharP.neg_one_pow_char R' p + · exact neg_one_pow_char R' p /-- The `n`th power of a quadratic character is the trivial character, when `n` is even. -/ theorem IsQuadratic.pow_even {χ : MulChar R R'} (hχ : χ.IsQuadratic) {n : ℕ} (hn : Even n) : diff --git a/Mathlib/RingTheory/FiniteType.lean b/Mathlib/RingTheory/FiniteType.lean index 97c79c644e56a..551f8541e5e2a 100644 --- a/Mathlib/RingTheory/FiniteType.lean +++ b/Mathlib/RingTheory/FiniteType.lean @@ -3,6 +3,7 @@ Copyright (c) 2020 Johan Commelin. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin -/ +import Mathlib.Algebra.FreeAlgebra import Mathlib.RingTheory.Adjoin.Tower import Mathlib.RingTheory.Ideal.QuotientOperations diff --git a/Mathlib/RingTheory/Perfection.lean b/Mathlib/RingTheory/Perfection.lean index 38d57045aa3eb..4206fc2c57b98 100644 --- a/Mathlib/RingTheory/Perfection.lean +++ b/Mathlib/RingTheory/Perfection.lean @@ -3,6 +3,7 @@ Copyright (c) 2020 Kenny Lau. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau -/ +import Mathlib.Algebra.CharP.ExpChar import Mathlib.Algebra.CharP.Pi import Mathlib.Algebra.CharP.Quotient import Mathlib.Algebra.CharP.Subring diff --git a/Mathlib/RingTheory/Polynomial/Basic.lean b/Mathlib/RingTheory/Polynomial/Basic.lean index 16082eba5ec13..4a7e163739c83 100644 --- a/Mathlib/RingTheory/Polynomial/Basic.lean +++ b/Mathlib/RingTheory/Polynomial/Basic.lean @@ -3,7 +3,8 @@ Copyright (c) 2019 Kenny Lau. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau -/ -import Mathlib.Algebra.CharP.ExpChar +import Mathlib.Algebra.CharP.Defs +import Mathlib.Algebra.GeomSum import Mathlib.Algebra.MvPolynomial.CommRing import Mathlib.Algebra.MvPolynomial.Equiv import Mathlib.RingTheory.Polynomial.Content diff --git a/Mathlib/RingTheory/Polynomial/SeparableDegree.lean b/Mathlib/RingTheory/Polynomial/SeparableDegree.lean index 6d5e04ba59670..dbe6bd9a2047e 100644 --- a/Mathlib/RingTheory/Polynomial/SeparableDegree.lean +++ b/Mathlib/RingTheory/Polynomial/SeparableDegree.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jakob Scholbach -/ import Mathlib.Algebra.Algebra.Defs -import Mathlib.Algebra.CharP.ExpChar import Mathlib.FieldTheory.Separable /-! From 034f3de5106c6a0cba767eb01f6a5a443ad5ea20 Mon Sep 17 00:00:00 2001 From: Bryan Gin-ge Chen <49699333+dependabot[bot]@users.noreply.github.com> Date: Mon, 21 Oct 2024 23:17:31 +0000 Subject: [PATCH 251/505] chore(deps): dependabot updates (#18026) Doing this manually because dependabot doesn't know how to update build.in.yml. Co-authored-by: Bryan Gin-ge Chen --- .github/build.in.yml | 4 ++-- .github/workflows/bors.yml | 4 ++-- .github/workflows/bot_fix_style_comment.yaml | 2 +- .github/workflows/bot_fix_style_review.yaml | 2 +- .github/workflows/bot_fix_style_review_comment.yaml | 2 +- .github/workflows/build.yml | 4 ++-- .github/workflows/build_fork.yml | 4 ++-- .github/workflows/dependent-issues.yml | 2 +- .github/workflows/labels_from_comment.yml | 2 +- .github/workflows/lean4checker.yml | 2 +- .github/workflows/lint_and_suggest_pr.yml | 2 +- .github/workflows/maintainer_merge_comment.yml | 2 +- .github/workflows/maintainer_merge_review.yml | 2 +- .github/workflows/maintainer_merge_review_comment.yml | 2 +- .github/workflows/nolints.yml | 2 +- .github/workflows/update_dependencies.yml | 2 +- 16 files changed, 20 insertions(+), 20 deletions(-) diff --git a/.github/build.in.yml b/.github/build.in.yml index 3b635701d4469..4d0134b8b6c39 100644 --- a/.github/build.in.yml +++ b/.github/build.in.yml @@ -33,7 +33,7 @@ jobs: uses: credfeto/action-case-checker@v1.3.0 - name: Look for ignored files - uses: credfeto/action-no-ignored-files@v1.1.0 + uses: credfeto/action-no-ignored-files@v1.2.0 - name: "Check for Lean files with the executable bit set" shell: bash @@ -92,7 +92,7 @@ jobs: # The Hoskinson runners may not have jq installed, so do that now. - name: 'Setup jq' - uses: dcarbone/install-jq-action@v1.0.1 + uses: dcarbone/install-jq-action@v2.1.0 - name: install elan run: | diff --git a/.github/workflows/bors.yml b/.github/workflows/bors.yml index 4b91df5b00120..816f0864dc82c 100644 --- a/.github/workflows/bors.yml +++ b/.github/workflows/bors.yml @@ -43,7 +43,7 @@ jobs: uses: credfeto/action-case-checker@v1.3.0 - name: Look for ignored files - uses: credfeto/action-no-ignored-files@v1.1.0 + uses: credfeto/action-no-ignored-files@v1.2.0 - name: "Check for Lean files with the executable bit set" shell: bash @@ -102,7 +102,7 @@ jobs: # The Hoskinson runners may not have jq installed, so do that now. - name: 'Setup jq' - uses: dcarbone/install-jq-action@v1.0.1 + uses: dcarbone/install-jq-action@v2.1.0 - name: install elan run: | diff --git a/.github/workflows/bot_fix_style_comment.yaml b/.github/workflows/bot_fix_style_comment.yaml index 4414eb4b576b5..5e0cbbbffd648 100644 --- a/.github/workflows/bot_fix_style_comment.yaml +++ b/.github/workflows/bot_fix_style_comment.yaml @@ -41,7 +41,7 @@ jobs: - name: install Python if: steps.user_permission.outputs.require-result == 'true' - uses: actions/setup-python@v4 + uses: actions/setup-python@v5 with: python-version: 3.8 diff --git a/.github/workflows/bot_fix_style_review.yaml b/.github/workflows/bot_fix_style_review.yaml index 99fa4ae9a7dab..ec389af5ce65b 100644 --- a/.github/workflows/bot_fix_style_review.yaml +++ b/.github/workflows/bot_fix_style_review.yaml @@ -47,7 +47,7 @@ jobs: - name: install Python if: steps.user_permission.outputs.require-result == 'true' - uses: actions/setup-python@v4 + uses: actions/setup-python@v5 with: python-version: 3.8 diff --git a/.github/workflows/bot_fix_style_review_comment.yaml b/.github/workflows/bot_fix_style_review_comment.yaml index 097b24ba35954..d997a6ad7989a 100644 --- a/.github/workflows/bot_fix_style_review_comment.yaml +++ b/.github/workflows/bot_fix_style_review_comment.yaml @@ -45,7 +45,7 @@ jobs: - name: install Python if: steps.user_permission.outputs.require-result == 'true' - uses: actions/setup-python@v4 + uses: actions/setup-python@v5 with: python-version: 3.8 diff --git a/.github/workflows/build.yml b/.github/workflows/build.yml index d1e5048846e07..f1c9dfc8fc360 100644 --- a/.github/workflows/build.yml +++ b/.github/workflows/build.yml @@ -50,7 +50,7 @@ jobs: uses: credfeto/action-case-checker@v1.3.0 - name: Look for ignored files - uses: credfeto/action-no-ignored-files@v1.1.0 + uses: credfeto/action-no-ignored-files@v1.2.0 - name: "Check for Lean files with the executable bit set" shell: bash @@ -109,7 +109,7 @@ jobs: # The Hoskinson runners may not have jq installed, so do that now. - name: 'Setup jq' - uses: dcarbone/install-jq-action@v1.0.1 + uses: dcarbone/install-jq-action@v2.1.0 - name: install elan run: | diff --git a/.github/workflows/build_fork.yml b/.github/workflows/build_fork.yml index e3a3e5d64931d..b1bffcfcd5e0b 100644 --- a/.github/workflows/build_fork.yml +++ b/.github/workflows/build_fork.yml @@ -47,7 +47,7 @@ jobs: uses: credfeto/action-case-checker@v1.3.0 - name: Look for ignored files - uses: credfeto/action-no-ignored-files@v1.1.0 + uses: credfeto/action-no-ignored-files@v1.2.0 - name: "Check for Lean files with the executable bit set" shell: bash @@ -106,7 +106,7 @@ jobs: # The Hoskinson runners may not have jq installed, so do that now. - name: 'Setup jq' - uses: dcarbone/install-jq-action@v1.0.1 + uses: dcarbone/install-jq-action@v2.1.0 - name: install elan run: | diff --git a/.github/workflows/dependent-issues.yml b/.github/workflows/dependent-issues.yml index 9e0b3d2da4b71..16c75acafbf69 100644 --- a/.github/workflows/dependent-issues.yml +++ b/.github/workflows/dependent-issues.yml @@ -11,7 +11,7 @@ jobs: runs-on: ubuntu-latest # timeout-minutes: 3 steps: - - uses: styfle/cancel-workflow-action@0.11.0 + - uses: styfle/cancel-workflow-action@0.12.1 with: all_but_latest: true access_token: ${{ github.token }} diff --git a/.github/workflows/labels_from_comment.yml b/.github/workflows/labels_from_comment.yml index b40b8d4c4daf6..02e881f3e822c 100644 --- a/.github/workflows/labels_from_comment.yml +++ b/.github/workflows/labels_from_comment.yml @@ -15,7 +15,7 @@ jobs: steps: - name: Add label based on comment - uses: actions/github-script@v6 + uses: actions/github-script@v7 with: github-token: ${{ secrets.GITHUB_TOKEN }} script: | diff --git a/.github/workflows/lean4checker.yml b/.github/workflows/lean4checker.yml index ab862ac1b44f4..b98c03bba301e 100644 --- a/.github/workflows/lean4checker.yml +++ b/.github/workflows/lean4checker.yml @@ -25,7 +25,7 @@ jobs: # The Hoskinson runners may not have jq installed, so do that now. - name: 'Setup jq' - uses: dcarbone/install-jq-action@v1.0.1 + uses: dcarbone/install-jq-action@v2.1.0 - name: install elan run: | diff --git a/.github/workflows/lint_and_suggest_pr.yml b/.github/workflows/lint_and_suggest_pr.yml index 409e78ebc223e..47029833459db 100644 --- a/.github/workflows/lint_and_suggest_pr.yml +++ b/.github/workflows/lint_and_suggest_pr.yml @@ -17,7 +17,7 @@ jobs: - uses: actions/checkout@v4 - name: install Python - uses: actions/setup-python@v4 + uses: actions/setup-python@v5 with: python-version: 3.8 diff --git a/.github/workflows/maintainer_merge_comment.yml b/.github/workflows/maintainer_merge_comment.yml index db769c6592ab6..7cdc17fe4ab74 100644 --- a/.github/workflows/maintainer_merge_comment.yml +++ b/.github/workflows/maintainer_merge_comment.yml @@ -52,7 +52,7 @@ jobs: body: | 🚀 Pull request has been placed on the maintainer queue by ${{ github.event.comment.user.login }}. - name: Add label to PR - uses: actions/github-script@v6 + uses: actions/github-script@v7 with: github-token: ${{secrets.GITHUB_TOKEN}} script: | diff --git a/.github/workflows/maintainer_merge_review.yml b/.github/workflows/maintainer_merge_review.yml index c7d8357bbfa64..e3b5954d6486c 100644 --- a/.github/workflows/maintainer_merge_review.yml +++ b/.github/workflows/maintainer_merge_review.yml @@ -51,7 +51,7 @@ jobs: body: | 🚀 Pull request has been placed on the maintainer queue by ${{ github.event.review.user.login }}. - name: Add label to PR - uses: actions/github-script@v6 + uses: actions/github-script@v7 with: github-token: ${{secrets.GITHUB_TOKEN}} script: | diff --git a/.github/workflows/maintainer_merge_review_comment.yml b/.github/workflows/maintainer_merge_review_comment.yml index 79a5ea176d9d9..10cb9570bac49 100644 --- a/.github/workflows/maintainer_merge_review_comment.yml +++ b/.github/workflows/maintainer_merge_review_comment.yml @@ -50,7 +50,7 @@ jobs: body: | 🚀 Pull request has been placed on the maintainer queue by ${{ github.event.comment.user.login }}. - name: Add label to PR - uses: actions/github-script@v6 + uses: actions/github-script@v7 with: github-token: ${{secrets.GITHUB_TOKEN}} script: | diff --git a/.github/workflows/nolints.yml b/.github/workflows/nolints.yml index 847e24854e855..33fda1bb1154c 100644 --- a/.github/workflows/nolints.yml +++ b/.github/workflows/nolints.yml @@ -16,7 +16,7 @@ jobs: # The Hoskinson runners may not have jq installed, so do that now. - name: 'Setup jq' - uses: dcarbone/install-jq-action@v1.0.1 + uses: dcarbone/install-jq-action@v2.1.0 - name: install elan run: | diff --git a/.github/workflows/update_dependencies.yml b/.github/workflows/update_dependencies.yml index b415c88c6a8c8..27646aaec561b 100644 --- a/.github/workflows/update_dependencies.yml +++ b/.github/workflows/update_dependencies.yml @@ -57,7 +57,7 @@ jobs: - name: Create Pull Request if: ${{ !contains(steps.PR.outputs.pr_labels, 'ready-to-merge') }} - uses: peter-evans/create-pull-request@v6 + uses: peter-evans/create-pull-request@v7 with: token: "${{ secrets.UPDATE_DEPENDENCIES_TOKEN }}" commit-message: "chore: update Mathlib dependencies ${{ env.timestamp }}" From cbf4f2267dd45537060bf4ef93c6bac6393c2c40 Mon Sep 17 00:00:00 2001 From: mathlib4-update-dependencies-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Tue, 22 Oct 2024 01:04:14 +0000 Subject: [PATCH 252/505] chore: update Mathlib dependencies 2024-10-22 (#18035) This PR updates the Mathlib dependencies. --- lake-manifest.json | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lake-manifest.json b/lake-manifest.json index 2736602d58cb5..89a2da4e5b414 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,7 +5,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "dd6b1019b5cef990161bf3edfebeb6b0be78044a", + "rev": "7c5548eeeb1748da5c2872dbd866d3acbaea4b3f", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", From fa9f767f542d37d4b39ac10a2863c8367b67fce0 Mon Sep 17 00:00:00 2001 From: Yakov Pechersky Date: Tue, 22 Oct 2024 04:01:53 +0000 Subject: [PATCH 253/505] chore(Topology/Algebra/Valued/NormedValued): rely on TC IsUltrametricDist instead of explicit arg (#17561) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Before, we had ``` [hK : NormedField K] (h : IsNonarchimedean (norm : K → ℝ)) ``` Now we use `[hK : NormedField K] [IsUltrametricDist K]` to convert between ultrametrically normed fields and valued fields. This allows the instance to be a scoped instance instead of a def, since `IsUltrametricDist` is a class on the type, as opposed to a plain prop about the norm of the type. The instance is scoped because otherwise, one gets a TC loop. --- .../Topology/Algebra/Valued/NormedValued.lean | 44 ++++++++++++++++--- 1 file changed, 39 insertions(+), 5 deletions(-) diff --git a/Mathlib/Topology/Algebra/Valued/NormedValued.lean b/Mathlib/Topology/Algebra/Valued/NormedValued.lean index 2c9ce7c08f2c8..249d5250b9868 100644 --- a/Mathlib/Topology/Algebra/Valued/NormedValued.lean +++ b/Mathlib/Topology/Algebra/Valued/NormedValued.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: María Inés de Frutos-Fernández -/ import Mathlib.Analysis.Normed.Field.Basic -import Mathlib.Analysis.Normed.Group.Uniform +import Mathlib.Analysis.Normed.Group.Ultra import Mathlib.RingTheory.Valuation.RankOne import Mathlib.Topology.Algebra.Valued.ValuationTopology @@ -30,7 +30,9 @@ open Filter Set Valuation open scoped NNReal -variable {K : Type*} [hK : NormedField K] (h : IsNonarchimedean (norm : K → ℝ)) +section + +variable {K : Type*} [hK : NormedField K] [IsUltrametricDist K] namespace NormedField @@ -40,15 +42,16 @@ def valuation : Valuation K ℝ≥0 where map_zero' := nnnorm_zero map_one' := nnnorm_one map_mul' := nnnorm_mul - map_add_le_max' := h + map_add_le_max' := IsUltrametricDist.norm_add_le_max -theorem valuation_apply (x : K) : valuation h x = ‖x‖₊ := rfl +@[simp] +theorem valuation_apply (x : K) : valuation x = ‖x‖₊ := rfl /-- The valued field structure on a nonarchimedean normed field `K`, determined by the norm. -/ def toValued : Valued K ℝ≥0 := { hK.toUniformSpace, @NonUnitalNormedRing.toNormedAddCommGroup K _ with - v := valuation h + v := valuation is_topological_valuation := fun U => by rw [Metric.mem_nhds_iff] exact ⟨fun ⟨ε, hε, h⟩ => @@ -56,8 +59,18 @@ def toValued : Valued K ℝ≥0 := fun ⟨ε, hε⟩ => ⟨(ε : ℝ), NNReal.coe_pos.mpr (Units.zero_lt _), fun x hx ↦ hε (mem_ball_zero_iff.mp hx)⟩⟩ } +instance {K : Type*} [NontriviallyNormedField K] [IsUltrametricDist K] : + Valuation.RankOne (valuation (K := K)) where + hom := .id _ + strictMono' := strictMono_id + nontrivial' := (exists_one_lt_norm K).imp fun x h ↦ by + have h' : x ≠ 0 := norm_eq_zero.not.mp (h.gt.trans' (by simp)).ne' + simp [valuation_apply, ← NNReal.coe_inj, h.ne', h'] + end NormedField +end + namespace Valued variable {L : Type*} [Field L] {Γ₀ : Type*} [LinearOrderedCommGroupWithZero Γ₀] @@ -126,4 +139,25 @@ def toNormedField : NormedField L := exact ⟨fun a b hab => lt_of_lt_of_le hab (min_le_left _ _), fun a b hab => lt_of_lt_of_le hab (min_le_right _ _)⟩ } +-- When a field is valued, one inherits a `NormedField`. +-- Scoped instance to avoid a typeclass loop or non-defeq topology or norms. +scoped[Valued] attribute [instance] Valued.toNormedField +scoped[NormedField] attribute [instance] NormedField.toValued + +section NormedField + +open scoped Valued + +protected lemma isNonarchimedean_norm : IsNonarchimedean ((‖·‖): L → ℝ) := Valued.norm_add_le + +instance : IsUltrametricDist L := + ⟨fun x y z ↦ by + refine (Valued.norm_add_le (x - y) (y - z)).trans_eq' ?_ + simp only [sub_add_sub_cancel] + rfl ⟩ + +lemma coe_valuation_eq_rankOne_hom_comp_valuation : ⇑NormedField.valuation = hv.hom ∘ val.v := rfl + +end NormedField + end Valued From d50eac5610440ba3d59f8a4779889ba4ececb1af Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Tue, 22 Oct 2024 04:32:03 +0000 Subject: [PATCH 254/505] chore: remove CoeFun instances where FunLike is available (#17870) During the port we found that `FunLike` is robust enough not to need an extra `CoeFun` shortcut. Let's see if that rule can be consistently applied to the whole of the library. There is still duplication between `FunLike` and `CoeFun` for `Grp`, `Mon`, `CommGrp` and `CommMon`, which will need a more thorough fix. See also https://github.com/leanprover-community/mathlib4/pull/17866. There is also a shortcut instance that we need for speed reasons: `SemilinearIsometry.instCoeFun`. I don't know why this is so performance-sensitive but the benchmarks tell us that it is. See also #17871 for removing porting notes and comments where this change happened earlier on. --- Mathlib/Analysis/Normed/Operator/LinearIsometry.lean | 9 ++++++--- Mathlib/Combinatorics/Young/SemistandardTableau.lean | 4 ---- Mathlib/Data/DFinsupp/Basic.lean | 5 ----- Mathlib/Data/Finsupp/Defs.lean | 5 ----- 4 files changed, 6 insertions(+), 17 deletions(-) diff --git a/Mathlib/Analysis/Normed/Operator/LinearIsometry.lean b/Mathlib/Analysis/Normed/Operator/LinearIsometry.lean index d3192d513de52..0146ea7262fd7 100644 --- a/Mathlib/Analysis/Normed/Operator/LinearIsometry.lean +++ b/Mathlib/Analysis/Normed/Operator/LinearIsometry.lean @@ -487,9 +487,12 @@ instance instSemilinearIsometryEquivClass : map_smulₛₗ e := map_smulₛₗ e.toLinearEquiv norm_map e := e.norm_map' --- TODO: Shouldn't these `CoeFun` instances be scrapped? -/-- Helper instance for when there's too many metavariables to apply `DFunLike.hasCoeToFun` -directly. -/ +/-- Shortcut instance, saving 8.5% of compilation time in +`Mathlib.Analysis.InnerProductSpace.Adjoint`. + +(This instance was pinpointed by benchmarks; we didn't do an in depth investigation why it is +specifically needed.) +-/ instance instCoeFun : CoeFun (E ≃ₛₗᵢ[σ₁₂] E₂) fun _ ↦ E → E₂ := ⟨DFunLike.coe⟩ theorem coe_injective : @Function.Injective (E ≃ₛₗᵢ[σ₁₂] E₂) (E → E₂) (↑) := diff --git a/Mathlib/Combinatorics/Young/SemistandardTableau.lean b/Mathlib/Combinatorics/Young/SemistandardTableau.lean index 1f4c8542dcbca..fac50829baf05 100644 --- a/Mathlib/Combinatorics/Young/SemistandardTableau.lean +++ b/Mathlib/Combinatorics/Young/SemistandardTableau.lean @@ -68,10 +68,6 @@ instance instFunLike {μ : YoungDiagram} : FunLike (SemistandardYoungTableau μ) cases T' congr -/-- Helper instance for when there's too many metavariables to apply `CoeFun.coe` directly. -/ -instance {μ : YoungDiagram} : CoeFun (SemistandardYoungTableau μ) fun _ ↦ ℕ → ℕ → ℕ := - inferInstance - @[simp] theorem to_fun_eq_coe {μ : YoungDiagram} {T : SemistandardYoungTableau μ} : T.entry = (T : ℕ → ℕ → ℕ) := diff --git a/Mathlib/Data/DFinsupp/Basic.lean b/Mathlib/Data/DFinsupp/Basic.lean index e8b0e30b975ab..702ce8b5b5fab 100644 --- a/Mathlib/Data/DFinsupp/Basic.lean +++ b/Mathlib/Data/DFinsupp/Basic.lean @@ -81,11 +81,6 @@ instance instDFunLike : DFunLike (Π₀ i, β i) ι β := congr subsingleton ⟩ -/-- Helper instance for when there are too many metavariables to apply `DFunLike.coeFunForall` -directly. -/ -instance : CoeFun (Π₀ i, β i) fun _ => ∀ i, β i := - inferInstance - @[simp] theorem toFun_eq_coe (f : Π₀ i, β i) : f.toFun = f := rfl diff --git a/Mathlib/Data/Finsupp/Defs.lean b/Mathlib/Data/Finsupp/Defs.lean index 17844597d0e12..c424ca7128840 100644 --- a/Mathlib/Data/Finsupp/Defs.lean +++ b/Mathlib/Data/Finsupp/Defs.lean @@ -116,11 +116,6 @@ instance instFunLike : FunLike (α →₀ M) α M := ext a exact (hf _).trans (hg _).symm⟩ -/-- Helper instance for when there are too many metavariables to apply the `DFunLike` instance -directly. -/ -instance instCoeFun : CoeFun (α →₀ M) fun _ => α → M := - inferInstance - @[ext] theorem ext {f g : α →₀ M} (h : ∀ a, f a = g a) : f = g := DFunLike.ext _ _ h From 07013722cda97276225650bec2a9635f95807ed6 Mon Sep 17 00:00:00 2001 From: Salvatore Mercuri Date: Tue, 22 Oct 2024 07:59:06 +0000 Subject: [PATCH 255/505] feat: `WithAbs` type synonym (#17998) This PR contains the type synonym `WithAbs`. This is a type synonym for a semiring `R` that depends on an absolute value `v : AbsoluteValue R S`, where `S` is an `OrderedSemiring`. This is used in order to handle ambiguities arising from multiple sources of instances. For example, each infinite place on a number field defines distinct uniform structures via their associated real absolute values. The `WithAbs` type synonym allows the type class inferencing system to automatically infer such dependent instances, making it simpler to define the Archimedean completions of a number field. See #16483 for this application along with prior feedback. --- Mathlib.lean | 1 + Mathlib/Algebra/Ring/WithAbs.lean | 100 +++++++++++++++++++++++ Mathlib/Analysis/Normed/Field/Basic.lean | 27 ++++++ 3 files changed, 128 insertions(+) create mode 100644 Mathlib/Algebra/Ring/WithAbs.lean diff --git a/Mathlib.lean b/Mathlib.lean index 9ca592ec191d3..c9febbcc1197e 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -817,6 +817,7 @@ import Mathlib.Algebra.Ring.Subsemiring.Pointwise import Mathlib.Algebra.Ring.SumsOfSquares import Mathlib.Algebra.Ring.ULift import Mathlib.Algebra.Ring.Units +import Mathlib.Algebra.Ring.WithAbs import Mathlib.Algebra.Ring.WithZero import Mathlib.Algebra.RingQuot import Mathlib.Algebra.SMulWithZero diff --git a/Mathlib/Algebra/Ring/WithAbs.lean b/Mathlib/Algebra/Ring/WithAbs.lean new file mode 100644 index 0000000000000..3b3fcb16224a3 --- /dev/null +++ b/Mathlib/Algebra/Ring/WithAbs.lean @@ -0,0 +1,100 @@ +/- +Copyright (c) 2024 Salvatore Mercuri. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Salvatore Mercuri +-/ +import Mathlib.Algebra.Group.Basic +import Mathlib.Algebra.Order.AbsoluteValue +import Mathlib.Analysis.Normed.Field.Basic + +/-! +# WithAbs + +`WithAbs v` is a type synonym for a semiring `R` which depends on an absolute value. The point of +this is to allow the type class inference system to handle multiple sources of instances that +arise from absolute values. See `NumberTheory.NumberField.Completion` for an example of this +being used to define Archimedean completions of a number field. + +## Main definitions + - `WithAbs` : type synonym for a semiring which depends on an absolute value. This is + a function that takes an absolute value on a semiring and returns the semiring. This can be used + to assign and infer instances on a semiring that depend on absolute values. + - `WithAbs.equiv v` : the canonical (type) equivalence between `WithAbs v` and `R`. + - `WithAbs.ringEquiv v` : The canonical ring equivalence between `WithAbs v` and `R`. +-/ +noncomputable section + +variable {R S K : Type*} [Semiring R] [OrderedSemiring S] [Field K] + +/-- Type synonym for a semiring which depends on an absolute value. This is a function that takes +an absolute value on a semiring and returns the semiring. We use this to assign and infer instances +on a semiring that depend on absolute values. -/ +@[nolint unusedArguments] +def WithAbs : AbsoluteValue R S → Type _ := fun _ => R + +namespace WithAbs + +variable (v : AbsoluteValue R ℝ) + +/-- Canonical equivalence between `WithAbs v` and `R`. -/ +def equiv : WithAbs v ≃ R := Equiv.refl (WithAbs v) + +instance instNonTrivial [Nontrivial R] : Nontrivial (WithAbs v) := inferInstanceAs (Nontrivial R) + +instance instUnique [Unique R] : Unique (WithAbs v) := inferInstanceAs (Unique R) + +instance instSemiring : Semiring (WithAbs v) := inferInstanceAs (Semiring R) + +instance instRing [Ring R] : Ring (WithAbs v) := inferInstanceAs (Ring R) + +instance instInhabited : Inhabited (WithAbs v) := ⟨0⟩ + +instance normedRing {R : Type*} [Ring R] (v : AbsoluteValue R ℝ) : NormedRing (WithAbs v) := + v.toNormedRing + +instance normedField (v : AbsoluteValue K ℝ) : NormedField (WithAbs v) := + v.toNormedField + +/-! `WithAbs.equiv` preserves the ring structure. -/ + +variable (x y : WithAbs v) (r s : R) +@[simp] +theorem equiv_zero : WithAbs.equiv v 0 = 0 := rfl + +@[simp] +theorem equiv_symm_zero : (WithAbs.equiv v).symm 0 = 0 := rfl + +@[simp] +theorem equiv_add : WithAbs.equiv v (x + y) = WithAbs.equiv v x + WithAbs.equiv v y := rfl + +@[simp] +theorem equiv_symm_add : + (WithAbs.equiv v).symm (r + s) = (WithAbs.equiv v).symm r + (WithAbs.equiv v).symm s := + rfl + +@[simp] +theorem equiv_sub [Ring R] : WithAbs.equiv v (x - y) = WithAbs.equiv v x - WithAbs.equiv v y := rfl + +@[simp] +theorem equiv_symm_sub [Ring R] : + (WithAbs.equiv v).symm (r - s) = (WithAbs.equiv v).symm r - (WithAbs.equiv v).symm s := + rfl + +@[simp] +theorem equiv_neg [Ring R] : WithAbs.equiv v (-x) = - WithAbs.equiv v x := rfl + +@[simp] +theorem equiv_symm_neg [Ring R] : (WithAbs.equiv v).symm (-r) = - (WithAbs.equiv v).symm r := rfl + +@[simp] +theorem equiv_mul : WithAbs.equiv v (x * y) = WithAbs.equiv v x * WithAbs.equiv v y := rfl + +@[simp] +theorem equiv_symm_mul : + (WithAbs.equiv v).symm (x * y) = (WithAbs.equiv v).symm x * (WithAbs.equiv v).symm y := + rfl + +/-- `WithAbs.equiv` as a ring equivalence. -/ +def ringEquiv : WithAbs v ≃+* R := RingEquiv.refl _ + +end WithAbs diff --git a/Mathlib/Analysis/Normed/Field/Basic.lean b/Mathlib/Analysis/Normed/Field/Basic.lean index 4368eb6b8b5c3..d548024cd56e3 100644 --- a/Mathlib/Analysis/Normed/Field/Basic.lean +++ b/Mathlib/Analysis/Normed/Field/Basic.lean @@ -17,6 +17,11 @@ definitions. Some useful results that relate the topology of the normed field to the discrete topology include: * `norm_eq_one_iff_ne_zero_of_discrete` + +Methods for constructing a normed ring/field instance from a given real absolute value on a +ring/field are given in: +* AbsoluteValue.toNormedRing +* AbsoluteValue.toNormedField -/ -- Guard against import creep. @@ -1090,3 +1095,25 @@ instance toNormedCommRing [NormedCommRing R] [SubringClass S R] (s : S) : Normed { SubringClass.toNormedRing s with mul_comm := mul_comm } end SubringClass + +namespace AbsoluteValue + +/-- A real absolute value on a ring determines a `NormedRing` structure. -/ +noncomputable def toNormedRing {R : Type*} [Ring R] (v : AbsoluteValue R ℝ) : NormedRing R where + norm := v + dist_eq _ _ := rfl + dist_self x := by simp only [sub_self, MulHom.toFun_eq_coe, AbsoluteValue.coe_toMulHom, map_zero] + dist_comm := v.map_sub + dist_triangle := v.sub_le + edist_dist x y := rfl + norm_mul x y := (v.map_mul x y).le + eq_of_dist_eq_zero := by simp only [MulHom.toFun_eq_coe, AbsoluteValue.coe_toMulHom, + AbsoluteValue.map_sub_eq_zero_iff, imp_self, implies_true] + +/-- A real absolute value on a field determines a `NormedField` structure. -/ +noncomputable def toNormedField {K : Type*} [Field K] (v : AbsoluteValue K ℝ) : NormedField K where + toField := inferInstanceAs (Field K) + __ := v.toNormedRing + norm_mul' := v.map_mul + +end AbsoluteValue From 970a6567fb2a41652d0ef34ee38c464dc2b76233 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Tue, 22 Oct 2024 08:25:18 +0000 Subject: [PATCH 256/505] feat(AlgebraicGeometry): basic properties of separated morphisms and schemes (#17960) --- Mathlib.lean | 1 + .../AlgebraicGeometry/Morphisms/Affine.lean | 2 +- .../Morphisms/ClosedImmersion.lean | 10 ++- .../Morphisms/Separated.lean | 79 ++++++++++++++++++- .../Limits/Constructions/BinaryProducts.lean | 40 ++++++++++ .../Limits/Shapes/Pullback/Equalizer.lean | 48 +++++++++++ 6 files changed, 173 insertions(+), 7 deletions(-) create mode 100644 Mathlib/CategoryTheory/Limits/Shapes/Pullback/Equalizer.lean diff --git a/Mathlib.lean b/Mathlib.lean index c9febbcc1197e..0ec50201451fc 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1713,6 +1713,7 @@ import Mathlib.CategoryTheory.Limits.Shapes.Products import Mathlib.CategoryTheory.Limits.Shapes.Pullback.Assoc import Mathlib.CategoryTheory.Limits.Shapes.Pullback.CommSq import Mathlib.CategoryTheory.Limits.Shapes.Pullback.Cospan +import Mathlib.CategoryTheory.Limits.Shapes.Pullback.Equalizer import Mathlib.CategoryTheory.Limits.Shapes.Pullback.HasPullback import Mathlib.CategoryTheory.Limits.Shapes.Pullback.Iso import Mathlib.CategoryTheory.Limits.Shapes.Pullback.Mono diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Affine.lean b/Mathlib/AlgebraicGeometry/Morphisms/Affine.lean index bcb0509ca655d..10ec1471d5adb 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Affine.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Affine.lean @@ -45,7 +45,7 @@ the preimage of any affine open subset of `Y` is affine. -/ class IsAffineHom {X Y : Scheme} (f : X ⟶ Y) : Prop where isAffine_preimage : ∀ U : Y.Opens, IsAffineOpen U → IsAffineOpen (f ⁻¹ᵁ U) -lemma isAffineOpen.preimage {X Y : Scheme} {U : Y.Opens} (hU : IsAffineOpen U) +lemma IsAffineOpen.preimage {X Y : Scheme} {U : Y.Opens} (hU : IsAffineOpen U) (f : X ⟶ Y) [IsAffineHom f] : IsAffineOpen (f ⁻¹ᵁ U) := IsAffineHom.isAffine_preimage _ hU diff --git a/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean b/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean index 741e9d4b1146b..e6170d7e2cc33 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean @@ -120,8 +120,12 @@ lemma of_surjective_of_isAffine {X Y : Scheme} [IsAffine X] [IsAffine Y] (f : X apply spec_of_surjective exact h -/-- If `f ≫ g` is a closed immersion, then `f` is a closed immersion. -/ -theorem of_comp {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) [IsClosedImmersion g] +/-- +If `f ≫ g` and `g` are closed immersions, then `f` is a closed immersion. +Also see `IsClosedImmersion.of_comp` for the general version +where `g` is only required to be separated. +-/ +theorem of_comp_isClosedImmersion {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) [IsClosedImmersion g] [IsClosedImmersion (f ≫ g)] : IsClosedImmersion f where base_closed := by have h := isClosedEmbedding (f ≫ g) @@ -243,7 +247,7 @@ theorem isAffine_surjective_of_isAffine [IsClosedImmersion f] : rw [← affineTargetImageFactorization_comp f] at i ⊢ haveI := of_surjective_of_isAffine (affineTargetImageInclusion f) (affineTargetImageInclusion_app_surjective f) - haveI := IsClosedImmersion.of_comp (affineTargetImageFactorization f) + haveI := IsClosedImmersion.of_comp_isClosedImmersion (affineTargetImageFactorization f) (affineTargetImageInclusion f) haveI := isIso_of_injective_of_isAffine (affineTargetImageFactorization_app_injective f) exact ⟨isAffine_of_isIso (affineTargetImageFactorization f), diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Separated.lean b/Mathlib/AlgebraicGeometry/Morphisms/Separated.lean index a1d81cfa22536..7f44de364381e 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Separated.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Separated.lean @@ -1,12 +1,13 @@ /- Copyright (c) 2024 Christian Merten. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Christian Merten +Authors: Christian Merten, Andrew Yang -/ import Mathlib.AlgebraicGeometry.Morphisms.ClosedImmersion import Mathlib.AlgebraicGeometry.Morphisms.QuasiSeparated import Mathlib.AlgebraicGeometry.Pullbacks import Mathlib.CategoryTheory.MorphismProperty.Limits +import Mathlib.CategoryTheory.Limits.Shapes.Pullback.Equalizer /-! @@ -14,6 +15,11 @@ import Mathlib.CategoryTheory.MorphismProperty.Limits A morphism of schemes is separated if its diagonal morphism is a closed immmersion. +## Main definitions +- `AlgebraicGeometry.IsSeparated`: The class of separated morphisms. +- `AlgebraicGeometry.Scheme.IsSeparated`: The class of separated schemes. +- `AlgebraicGeometry.IsSeparated.hasAffineProperty`: + A morphism is separated iff the preimage of affine opens are separated schemes. -/ @@ -27,7 +33,7 @@ open scoped AlgebraicGeometry namespace AlgebraicGeometry -variable {X Y : Scheme.{u}} (f : X ⟶ Y) +variable {X Y Z : Scheme.{u}} (f : X ⟶ Y) (g : Y ⟶ Z) /-- A morphism is separated if the diagonal map is a closed immersion. -/ @[mk_iff] @@ -57,7 +63,7 @@ instance stableUnderComposition : MorphismProperty.IsStableUnderComposition @IsS rw [isSeparated_eq_diagonal_isClosedImmersion] exact MorphismProperty.diagonal_isStableUnderComposition IsClosedImmersion.stableUnderBaseChange -instance {Z : Scheme.{u}} (g : Y ⟶ Z) [IsSeparated f] [IsSeparated g] : IsSeparated (f ≫ g) := +instance [IsSeparated f] [IsSeparated g] : IsSeparated (f ≫ g) := stableUnderComposition.comp_mem f g inferInstance inferInstance instance : MorphismProperty.IsMultiplicative @IsSeparated where @@ -80,6 +86,7 @@ instance (R S : CommRingCat.{u}) (f : R ⟶ S) : IsSeparated (Spec.map f) := by (Algebra.TensorProduct.lmul'_apply_tmul (R := R) (S := S) 1 x).trans (one_mul x)⟩ instance (priority := 100) [h : IsAffineHom f] : IsSeparated f := by + clear g Z wlog hY : IsAffine Y · rw [IsLocalAtTarget.iff_of_iSup_eq_top (P := @IsSeparated) _ (iSup_affineOpens_eq_top Y)] @@ -90,6 +97,72 @@ instance (priority := 100) [h : IsAffineHom f] : IsSeparated f := by rw [MorphismProperty.arrow_mk_iso_iff @IsSeparated (arrowIsoSpecΓOfIsAffine f)] infer_instance +instance {S T : Scheme.{u}} (f : X ⟶ S) (g : Y ⟶ S) (i : S ⟶ T) [IsSeparated i] : + IsClosedImmersion (pullback.mapDesc f g i) := + IsClosedImmersion.stableUnderBaseChange (pullback_map_diagonal_isPullback f g i) + inferInstance + +/-- Given `f : X ⟶ Y` and `g : Y ⟶ Z` such that `g` is separated, the induced map +`X ⟶ X ×[Z] Y` is a closed immersion. -/ +instance [IsSeparated g] : + IsClosedImmersion (pullback.lift (𝟙 _) f (Category.id_comp (f ≫ g))) := by + rw [← MorphismProperty.cancel_left_of_respectsIso @IsClosedImmersion (pullback.fst f (𝟙 Y))] + rw [← MorphismProperty.cancel_right_of_respectsIso @IsClosedImmersion _ + (pullback.congrHom rfl (Category.id_comp g)).inv] + convert (inferInstanceAs <| IsClosedImmersion (pullback.mapDesc f (𝟙 _) g)) using 1 + ext : 1 <;> simp [pullback.condition] + end IsSeparated +lemma IsClosedImmersion.of_comp [IsClosedImmersion (f ≫ g)] [IsSeparated g] : + IsClosedImmersion f := by + rw [← pullback.lift_snd (𝟙 _) f (Category.id_comp (f ≫ g))] + have := IsClosedImmersion.stableUnderBaseChange.snd (f ≫ g) g inferInstance + infer_instance + +lemma IsSeparated.of_comp [IsSeparated (f ≫ g)] : IsSeparated f := by + have := IsSeparated.diagonal_isClosedImmersion (f := f ≫ g) + rw [pullback.diagonal_comp] at this + exact ⟨@IsClosedImmersion.of_comp _ _ _ _ _ this inferInstance⟩ + +lemma IsSeparated.comp_iff [IsSeparated g] : IsSeparated (f ≫ g) ↔ IsSeparated f := + ⟨fun _ ↦ .of_comp f g, fun _ ↦ inferInstance⟩ + +namespace Scheme + +/-- A scheme `X` is separated if it is separated over `⊤_ Scheme`. -/ +@[mk_iff] +protected class IsSeparated (X : Scheme.{u}) : Prop where + isSeparated_terminal_from : IsSeparated (terminal.from X) + +attribute [instance] IsSeparated.isSeparated_terminal_from + +lemma isSeparated_iff_isClosedImmersion_prod_lift {X : Scheme.{u}} : + X.IsSeparated ↔ IsClosedImmersion (prod.lift (𝟙 X) (𝟙 X)) := by + rw [isSeparated_iff, AlgebraicGeometry.isSeparated_iff, iff_iff_eq, + ← MorphismProperty.cancel_right_of_respectsIso @IsClosedImmersion _ (prodIsoPullback X X).hom] + congr + ext : 1 <;> simp + +instance [X.IsSeparated] : IsClosedImmersion (prod.lift (𝟙 X) (𝟙 X)) := by + rwa [← isSeparated_iff_isClosedImmersion_prod_lift] + +instance (priority := 900) {X : Scheme.{u}} [IsAffine X] : X.IsSeparated := ⟨inferInstance⟩ + +instance (priority := 900) [X.IsSeparated] : IsSeparated f := by + apply (config := { allowSynthFailures := true }) @IsSeparated.of_comp (g := terminal.from Y) + rw [terminal.comp_from] + infer_instance + +instance (f g : X ⟶ Y) [Y.IsSeparated] : IsClosedImmersion (Limits.equalizer.ι f g) := + IsClosedImmersion.stableUnderBaseChange (isPullback_equalizer_prod f g).flip inferInstance + +end Scheme + +instance IsSeparated.hasAffineProperty : + HasAffineProperty @IsSeparated fun X _ _ _ ↦ X.IsSeparated := by + convert HasAffineProperty.of_isLocalAtTarget @IsSeparated with X Y f hY + rw [Scheme.isSeparated_iff, ← terminal.comp_from f, IsSeparated.comp_iff] + rfl + end AlgebraicGeometry diff --git a/Mathlib/CategoryTheory/Limits/Constructions/BinaryProducts.lean b/Mathlib/CategoryTheory/Limits/Constructions/BinaryProducts.lean index 93baeb3c4ad9d..3278a8a966db7 100644 --- a/Mathlib/CategoryTheory/Limits/Constructions/BinaryProducts.lean +++ b/Mathlib/CategoryTheory/Limits/Constructions/BinaryProducts.lean @@ -109,6 +109,26 @@ noncomputable def prodIsoPullback [HasTerminal C] [HasPullbacks C] (X Y : C) [HasBinaryProduct X Y] : X ⨯ Y ≅ pullback (terminal.from X) (terminal.from Y) := limit.isoLimitCone (limitConeOfTerminalAndPullbacks _) +@[reassoc (attr := simp)] +lemma prodIsoPullback_hom_fst [HasTerminal C] [HasPullbacks C] (X Y : C) + [HasBinaryProduct X Y] : (prodIsoPullback X Y).hom ≫ pullback.fst _ _ = prod.fst := + limit.isoLimitCone_hom_π (limitConeOfTerminalAndPullbacks _) ⟨.left⟩ + +@[reassoc (attr := simp)] +lemma prodIsoPullback_hom_snd [HasTerminal C] [HasPullbacks C] (X Y : C) + [HasBinaryProduct X Y] : (prodIsoPullback X Y).hom ≫ pullback.snd _ _ = prod.snd := + limit.isoLimitCone_hom_π (limitConeOfTerminalAndPullbacks _) ⟨.right⟩ + +@[reassoc (attr := simp)] +lemma prodIsoPullback_inv_fst [HasTerminal C] [HasPullbacks C] (X Y : C) + [HasBinaryProduct X Y] : (prodIsoPullback X Y).inv ≫ prod.fst = pullback.fst _ _ := + limit.isoLimitCone_inv_π (limitConeOfTerminalAndPullbacks _) ⟨.left⟩ + +@[reassoc (attr := simp)] +lemma prodIsoPullback_inv_snd [HasTerminal C] [HasPullbacks C] (X Y : C) + [HasBinaryProduct X Y] : (prodIsoPullback X Y).inv ≫ prod.snd = pullback.snd _ _ := + limit.isoLimitCone_inv_π (limitConeOfTerminalAndPullbacks _) ⟨.right⟩ + /-- If a cospan is the pushout cospan under the initial object, then it is a binary coproduct. -/ def isBinaryCoproductOfIsInitialIsPushout (F : Discrete WalkingPair ⥤ C) (c : Cocone F) {X : C} (hX : IsInitial X) (f : X ⟶ F.obj ⟨WalkingPair.left⟩) (g : X ⟶ F.obj ⟨WalkingPair.right⟩) @@ -194,3 +214,23 @@ a coproduct of objects `X` and `Y` is isomorphic to a pushout. -/ noncomputable def coprodIsoPushout [HasInitial C] [HasPushouts C] (X Y : C) [HasBinaryCoproduct X Y] : X ⨿ Y ≅ pushout (initial.to X) (initial.to Y) := colimit.isoColimitCocone (colimitCoconeOfInitialAndPushouts _) + +@[reassoc (attr := simp)] +lemma inl_coprodIsoPushout_hom [HasInitial C] [HasPushouts C] (X Y : C) + [HasBinaryCoproduct X Y] : coprod.inl ≫ (coprodIsoPushout X Y).hom = pushout.inl _ _ := + colimit.isoColimitCocone_ι_hom (colimitCoconeOfInitialAndPushouts _) _ + +@[reassoc (attr := simp)] +lemma inr_coprodIsoPushout_hom [HasInitial C] [HasPushouts C] (X Y : C) + [HasBinaryCoproduct X Y] : coprod.inr ≫ (coprodIsoPushout X Y).hom = pushout.inr _ _ := + colimit.isoColimitCocone_ι_hom (colimitCoconeOfInitialAndPushouts _) _ + +@[reassoc (attr := simp)] +lemma inl_coprodIsoPushout_inv [HasInitial C] [HasPushouts C] (X Y : C) + [HasBinaryCoproduct X Y] : pushout.inl _ _ ≫ (coprodIsoPushout X Y).inv = coprod.inl := + colimit.isoColimitCocone_ι_inv (colimitCoconeOfInitialAndPushouts (pair X Y)) ⟨.left⟩ + +@[reassoc (attr := simp)] +lemma inr_coprodIsoPushout_inv [HasInitial C] [HasPushouts C] (X Y : C) + [HasBinaryCoproduct X Y] : pushout.inr _ _ ≫ (coprodIsoPushout X Y).inv = coprod.inr := + colimit.isoColimitCocone_ι_inv (colimitCoconeOfInitialAndPushouts (pair X Y)) ⟨.right⟩ diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Pullback/Equalizer.lean b/Mathlib/CategoryTheory/Limits/Shapes/Pullback/Equalizer.lean new file mode 100644 index 0000000000000..acd890c5e85aa --- /dev/null +++ b/Mathlib/CategoryTheory/Limits/Shapes/Pullback/Equalizer.lean @@ -0,0 +1,48 @@ +/- +Copyright (c) 2024 Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Andrew Yang +-/ +import Mathlib.CategoryTheory.Limits.Shapes.Pullback.CommSq + +/-! +# Equalizers as pullbacks of products + +Also see `CategoryTheory.Limits.Constructions.Equalizers` for very similar results. + +-/ + +universe v u + +namespace CategoryTheory.Limits + +variable {C : Type u} [Category.{v} C] {X Y : C} (f g : X ⟶ Y) + +/-- The equalizer of `f g : X ⟶ Y` is the pullback of the diagonal map `Y ⟶ Y × Y` +along the map `(f, g) : X ⟶ Y × Y`. -/ +lemma isPullback_equalizer_prod [HasEqualizer f g] [HasBinaryProduct Y Y] : + IsPullback (equalizer.ι f g) (equalizer.ι f g ≫ f) (prod.lift f g) (prod.lift (𝟙 _) (𝟙 _)) := by + refine ⟨⟨by ext <;> simp [equalizer.condition f g]⟩, ⟨PullbackCone.IsLimit.mk _ ?_ ?_ ?_ ?_⟩⟩ + · refine fun s ↦ equalizer.lift s.fst ?_ + have H₁ : s.fst ≫ f = s.snd := by simpa using congr($s.condition ≫ prod.fst) + have H₂ : s.fst ≫ g = s.snd := by simpa using congr($s.condition ≫ prod.snd) + exact H₁.trans H₂.symm + · exact fun s ↦ by simp + · exact fun s ↦ by simpa using congr($s.condition ≫ prod.fst) + · exact fun s m hm _ ↦ by ext; simp [*] + +/-- The coequalizer of `f g : X ⟶ Y` is the pushout of the diagonal map `X ⨿ X ⟶ X` +along the map `(f, g) : X ⨿ X ⟶ Y`. -/ +lemma isPushout_coequalizer_coprod [HasCoequalizer f g] [HasBinaryCoproduct X X] : + IsPushout (coprod.desc f g) (coprod.desc (𝟙 _) (𝟙 _)) + (coequalizer.π f g) (f ≫ coequalizer.π f g) := by + refine ⟨⟨by ext <;> simp [coequalizer.condition f g]⟩, ⟨PushoutCocone.IsColimit.mk _ ?_ ?_ ?_ ?_⟩⟩ + · refine fun s ↦ coequalizer.desc s.inl ?_ + have H₁ : f ≫ s.inl = s.inr := by simpa using congr(coprod.inl ≫ $s.condition) + have H₂ : g ≫ s.inl = s.inr := by simpa using congr(coprod.inr ≫ $s.condition) + exact H₁.trans H₂.symm + · exact fun s ↦ by simp + · exact fun s ↦ by simpa using congr(coprod.inl ≫ $s.condition) + · exact fun s m hm _ ↦ by ext; simp [*] + +end CategoryTheory.Limits From 25461520176cf70c9941f1a5a42cd736c69432b7 Mon Sep 17 00:00:00 2001 From: Peter Nelson <71660771+apnelson1@users.noreply.github.com> Date: Tue, 22 Oct 2024 08:37:11 +0000 Subject: [PATCH 257/505] feat(Data/Matroid/Closure): closure + insertion API (#15028) Some API for interactions between independence, closure and single-element insertions/differences in matroids. --- Mathlib/Data/Matroid/Closure.lean | 119 +++++++++++++++++++++++++++++- 1 file changed, 116 insertions(+), 3 deletions(-) diff --git a/Mathlib/Data/Matroid/Closure.lean b/Mathlib/Data/Matroid/Closure.lean index a788bf4ba0f3d..9849ae57bc974 100644 --- a/Mathlib/Data/Matroid/Closure.lean +++ b/Mathlib/Data/Matroid/Closure.lean @@ -69,7 +69,7 @@ the subtype `↑(Iic M.E)` via `Matroid.SubtypeClosure`, albeit less elegantly. open Set namespace Matroid -variable {ι α : Type*} {M : Matroid α} {F X Y : Set α} {e : α} +variable {ι α : Type*} {M : Matroid α} {F X Y : Set α} {e f : α} section Flat @@ -176,7 +176,7 @@ lemma closure_mono (M : Matroid α) : Monotone M.closure := @[simp] lemma closure_closure (M : Matroid α) (X : Set α) : M.closure (M.closure X) = M.closure X := (M.subset_closure _).antisymm' (subset_sInter - (fun _ hF ↦ (closure_subset_closure _ (sInter_subset_of_mem hF)).trans hF.1.closure.subset)) + (fun F hF ↦ (closure_subset_closure _ (sInter_subset_of_mem hF)).trans hF.1.closure.subset)) lemma closure_subset_closure_of_subset_closure (hXY : X ⊆ M.closure Y) : M.closure X ⊆ M.closure Y := @@ -274,7 +274,7 @@ lemma mem_closure_self (M : Matroid α) (e : α) (he : e ∈ M.E := by aesop_mat section Indep -variable {ι : Sort*} {I J B : Set α} {f x : α} +variable {ι : Sort*} {I J B : Set α} {x y : α} lemma Indep.closure_eq_setOf_basis_insert (hI : M.Indep I) : M.closure I = {x | M.Basis I (insert x I)} := by @@ -514,6 +514,119 @@ lemma Basis.eq_of_closure_subset (hI : M.Basis I X) (hJI : J ⊆ I) (hJ : X ⊆ @[simp] lemma empty_basis_iff : M.Basis ∅ X ↔ X ⊆ M.closure ∅ := by rw [basis_iff_indep_closure, and_iff_right M.empty_indep, and_iff_left (empty_subset _)] +lemma indep_iff_forall_not_mem_closure_diff (hI : I ⊆ M.E := by aesop_mat) : + M.Indep I ↔ ∀ ⦃e⦄, e ∈ I → e ∉ M.closure (I \ {e}) := by + use fun h e heI he ↦ ((h.closure_inter_eq_self_of_subset diff_subset).subset ⟨he, heI⟩).2 rfl + intro h + obtain ⟨J, hJ⟩ := M.exists_basis I + convert hJ.indep + refine hJ.subset.antisymm' (fun e he ↦ by_contra fun heJ ↦ h he ?_) + exact mem_of_mem_of_subset + (hJ.subset_closure he) (M.closure_subset_closure (subset_diff_singleton hJ.subset heJ)) + +/-- An alternative version of `Matroid.indep_iff_forall_not_mem_closure_diff` where the +hypothesis that `I ⊆ M.E` is contained in the RHS rather than the hypothesis. -/ +lemma indep_iff_forall_not_mem_closure_diff' : + M.Indep I ↔ I ⊆ M.E ∧ ∀ e ∈ I, e ∉ M.closure (I \ {e}) := + ⟨fun h ↦ ⟨h.subset_ground, (indep_iff_forall_not_mem_closure_diff h.subset_ground).mp h⟩, fun h ↦ + (indep_iff_forall_not_mem_closure_diff h.1).mpr h.2⟩ + +lemma Indep.not_mem_closure_diff_of_mem (hI : M.Indep I) (he : e ∈ I) : e ∉ M.closure (I \ {e}) := + (indep_iff_forall_not_mem_closure_diff'.1 hI).2 e he + +lemma indep_iff_forall_closure_diff_ne : + M.Indep I ↔ ∀ ⦃e⦄, e ∈ I → M.closure (I \ {e}) ≠ M.closure I := by + rw [indep_iff_forall_not_mem_closure_diff'] + refine ⟨fun ⟨hIE, h⟩ e heI h_eq ↦ h e heI (h_eq.symm.subset (M.mem_closure_of_mem heI)), + fun h ↦ ⟨fun e heI ↦ by_contra fun heE ↦ h heI ?_,fun e heI hin ↦ h heI ?_⟩⟩ + · rw [← closure_inter_ground, inter_comm, inter_diff_distrib_left, + inter_singleton_eq_empty.mpr heE, diff_empty, inter_comm, closure_inter_ground] + nth_rw 2 [show I = insert e (I \ {e}) by simp [heI]] + rw [← closure_insert_closure_eq_closure_insert, insert_eq_of_mem hin, closure_closure] + +lemma Indep.closure_ssubset_closure (hI : M.Indep I) (hJI : J ⊂ I) : M.closure J ⊂ M.closure I := by + obtain ⟨e, heI, heJ⟩ := exists_of_ssubset hJI + exact (M.closure_subset_closure hJI.subset).ssubset_of_not_subset fun hss ↦ heJ <| + (hI.closure_inter_eq_self_of_subset hJI.subset).subset ⟨hss (M.mem_closure_of_mem heI), heI⟩ + +lemma indep_iff_forall_closure_ssubset_of_ssubset (hI : I ⊆ M.E := by aesop_mat) : + M.Indep I ↔ ∀ ⦃J⦄, J ⊂ I → M.closure J ⊂ M.closure I := by + refine ⟨fun h _ ↦ h.closure_ssubset_closure, + fun h ↦ (indep_iff_forall_not_mem_closure_diff hI).2 fun e heI hecl ↦ ?_⟩ + refine (h (diff_singleton_sSubset.2 heI)).ne ?_ + rw [show I = insert e (I \ {e}) by simp [heI], ← closure_insert_closure_eq_closure_insert, + insert_eq_of_mem hecl] + simp + +lemma Indep.closure_diff_ssubset (hI : M.Indep I) (hX : (I ∩ X).Nonempty) : + M.closure (I \ X) ⊂ M.closure I := by + refine hI.closure_ssubset_closure <| diff_subset.ssubset_of_ne fun h ↦ ?_ + rw [sdiff_eq_left, disjoint_iff_inter_eq_empty] at h + simp [h] at hX + +lemma Indep.closure_diff_singleton_ssubset (hI : M.Indep I) (he : e ∈ I) : + M.closure (I \ {e}) ⊂ M.closure I := + hI.closure_ssubset_closure <| by simpa + end Indep +section insert + +lemma mem_closure_insert (he : e ∉ M.closure X) (hef : e ∈ M.closure (insert f X)) : + f ∈ M.closure (insert e X) := by + rw [← closure_inter_ground] at * + have hfE : f ∈ M.E := by + by_contra! hfE; rw [insert_inter_of_not_mem hfE] at hef; exact he hef + have heE : e ∈ M.E := (M.closure_subset_ground _) hef + rw [insert_inter_of_mem hfE] at hef; rw [insert_inter_of_mem heE] + + obtain ⟨I, hI⟩ := M.exists_basis (X ∩ M.E) + rw [← hI.closure_eq_closure, hI.indep.not_mem_closure_iff] at he + rw [← closure_insert_closure_eq_closure_insert, ← hI.closure_eq_closure, + closure_insert_closure_eq_closure_insert, he.1.mem_closure_iff] at * + rw [or_iff_not_imp_left, dep_iff, insert_comm, + and_iff_left (insert_subset heE (insert_subset hfE hI.indep.subset_ground)), not_not] + intro h + rw [(h.subset (subset_insert _ _)).mem_closure_iff, or_iff_right (h.not_dep), mem_insert_iff, + or_iff_left he.2] at hef + subst hef; apply mem_insert + +lemma closure_exchange (he : e ∈ M.closure (insert f X) \ M.closure X) : + f ∈ M.closure (insert e X) \ M.closure X := + ⟨mem_closure_insert he.2 he.1, fun hf ↦ by + rwa [closure_insert_eq_of_mem_closure hf, diff_self, iff_false_intro (not_mem_empty _)] at he⟩ + +lemma closure_exchange_iff : + e ∈ M.closure (insert f X) \ M.closure X ↔ f ∈ M.closure (insert e X) \ M.closure X := + ⟨closure_exchange, closure_exchange⟩ + +lemma closure_insert_congr (he : e ∈ M.closure (insert f X) \ M.closure X) : + M.closure (insert e X) = M.closure (insert f X) := by + have hf := closure_exchange he + rw [eq_comm, ← closure_closure, ← insert_eq_of_mem he.1, closure_insert_closure_eq_closure_insert, + insert_comm, ← closure_closure, ← closure_insert_closure_eq_closure_insert, + insert_eq_of_mem hf.1, closure_closure, closure_closure] + +lemma closure_diff_eq_self (h : Y ⊆ M.closure (X \ Y)) : M.closure (X \ Y) = M.closure X := by + rw [← diff_union_inter X Y, ← closure_union_closure_left_eq, + union_eq_self_of_subset_right (inter_subset_right.trans h), closure_closure, diff_union_inter] + +lemma closure_diff_singleton_eq_closure (h : e ∈ M.closure (X \ {e})) : + M.closure (X \ {e}) = M.closure X := + closure_diff_eq_self (by simpa) + +lemma subset_closure_diff_iff_closure_eq (h : Y ⊆ X) (hY : Y ⊆ M.E := by aesop_mat) : + Y ⊆ M.closure (X \ Y) ↔ M.closure (X \ Y) = M.closure X := + ⟨closure_diff_eq_self, fun h' ↦ (M.subset_closure_of_subset' h).trans h'.symm.subset⟩ + +lemma mem_closure_diff_singleton_iff_closure (he : e ∈ X) (heE : e ∈ M.E := by aesop_mat) : + e ∈ M.closure (X \ {e}) ↔ M.closure (X \ {e}) = M.closure X := by + simpa using subset_closure_diff_iff_closure_eq (Y := {e}) (X := X) (by simpa) + +end insert + +lemma ext_closure {M₁ M₂ : Matroid α} (h : ∀ X, M₁.closure X = M₂.closure X) : M₁ = M₂ := + eq_of_indep_iff_indep_forall (by simpa using h univ) + (fun _ _ ↦ by simp_rw [indep_iff_forall_closure_diff_ne, h]) + end Matroid From 1ece609f6ff5522cd5858de729b808e787a61695 Mon Sep 17 00:00:00 2001 From: Oliver Nash Date: Tue, 22 Oct 2024 08:37:13 +0000 Subject: [PATCH 258/505] feat: define finitely-semisimple linear maps (#17689) The motivation is to be able to apply `IsFinitelySemisimple.maxGenEigenspace_eq_eigenspace` to symmetric maps in infinite dimensions. Such maps are not in general semisimple but are semisimple on each invariant finite-dimensional subspace. --- Mathlib.lean | 1 + Mathlib/Algebra/Lie/Weights/Killing.lean | 5 +- .../Algebra/Module/Submodule/Invariant.lean | 119 +++++++++++++++ Mathlib/Algebra/Module/Submodule/Map.lean | 8 +- Mathlib/Algebra/Module/Submodule/Range.lean | 6 + Mathlib/Algebra/Polynomial/Module/AEval.lean | 101 ++++++------- .../LinearAlgebra/Eigenspace/Semisimple.lean | 16 +- Mathlib/LinearAlgebra/Semisimple.lean | 141 +++++++++++++++--- Mathlib/Order/Disjoint.lean | 4 + Mathlib/Order/Hom/Set.lean | 33 +++- Mathlib/Order/LatticeIntervals.lean | 10 ++ 11 files changed, 360 insertions(+), 84 deletions(-) create mode 100644 Mathlib/Algebra/Module/Submodule/Invariant.lean diff --git a/Mathlib.lean b/Mathlib.lean index 0ec50201451fc..01695686dd757 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -502,6 +502,7 @@ import Mathlib.Algebra.Module.Submodule.Basic import Mathlib.Algebra.Module.Submodule.Bilinear import Mathlib.Algebra.Module.Submodule.EqLocus import Mathlib.Algebra.Module.Submodule.Equiv +import Mathlib.Algebra.Module.Submodule.Invariant import Mathlib.Algebra.Module.Submodule.IterateMapComap import Mathlib.Algebra.Module.Submodule.Ker import Mathlib.Algebra.Module.Submodule.Lattice diff --git a/Mathlib/Algebra/Lie/Weights/Killing.lean b/Mathlib/Algebra/Lie/Weights/Killing.lean index 70df35ad7a117..3fe00e4149748 100644 --- a/Mathlib/Algebra/Lie/Weights/Killing.lean +++ b/Mathlib/Algebra/Lie/Weights/Killing.lean @@ -269,7 +269,7 @@ lemma isSemisimple_ad_of_mem_isCartanSubalgebra {x : L} (hx : x ∈ H) : (genWeightSpace_le_genWeightSpaceOf L x' α) hy rw [maxGenEigenspace_eq] at hy set k := maxGenEigenspaceIndex (ad K L x) (α x') - rw [apply_eq_of_mem_genEigenspace_of_comm_of_isSemisimple_of_isNilpotent_sub hy hS₀ hS hN] + rw [apply_eq_of_mem_of_comm_of_isFinitelySemisimple_of_isNil hy hS₀ hS.isFinitelySemisimple hN] /- So `S` obeys the derivation axiom if we restrict to root spaces. -/ have h_der (y z : L) (α β : H → K) (hy : y ∈ rootSpace H α) (hz : z ∈ rootSpace H β) : S ⁅y, z⁆ = ⁅S y, z⁆ + ⁅y, S z⁆ := by @@ -317,7 +317,8 @@ lemma lie_eq_smul_of_mem_rootSpace {α : H → K} {x : L} (hx : x ∈ rootSpace replace hx : x ∈ (ad K L h).maxGenEigenspace (α h) := genWeightSpace_le_genWeightSpaceOf L h α hx rw [(isSemisimple_ad_of_mem_isCartanSubalgebra - h.property).maxGenEigenspace_eq_eigenspace, Module.End.mem_eigenspace_iff] at hx + h.property).isFinitelySemisimple.maxGenEigenspace_eq_eigenspace, + Module.End.mem_eigenspace_iff] at hx simpa using hx lemma lie_eq_killingForm_smul_of_mem_rootSpace_of_mem_rootSpace_neg diff --git a/Mathlib/Algebra/Module/Submodule/Invariant.lean b/Mathlib/Algebra/Module/Submodule/Invariant.lean new file mode 100644 index 0000000000000..76e70e00cd1fa --- /dev/null +++ b/Mathlib/Algebra/Module/Submodule/Invariant.lean @@ -0,0 +1,119 @@ +/- +Copyright (c) 2024 Oliver Nash. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Oliver Nash +-/ +import Mathlib.Algebra.Module.Submodule.Map +import Mathlib.Order.Sublattice + +/-! +# The lattice of invariant submodules + +In this file we defined the type `Module.End.invtSubmodule`, associated to a linear endomorphism of +a module. Its utilty stems primarily from those occasions on which we wish to take advantage of the +lattice structure of invariant submodules. + +See also `Module.AEval`. + +-/ + +namespace Module.End + +variable {R M : Type*} [Semiring R] [AddCommMonoid M] [Module R M] (f : End R M) + +/-- Given an endomorphism, `f` of some module, this is the sublattice of all `f`-invariant +submodules. -/ +def invtSubmodule : Sublattice (Submodule R M) where + carrier := {p : Submodule R M | p ≤ p.comap f} + supClosed' p hp q hq := sup_le_iff.mpr + ⟨le_trans hp <| Submodule.comap_mono le_sup_left, + le_trans hq <| Submodule.comap_mono le_sup_right⟩ + infClosed' p hp q hq := by + simp only [Set.mem_setOf_eq, Submodule.comap_inf, le_inf_iff] + exact ⟨inf_le_of_left_le hp, inf_le_of_right_le hq⟩ + +lemma mem_invtSubmodule {p : Submodule R M} : + p ∈ f.invtSubmodule ↔ p ≤ p.comap f := + Iff.rfl + +namespace invtSubmodule + +@[simp] +protected lemma top_mem : ⊤ ∈ f.invtSubmodule := by simp [invtSubmodule] + +@[simp] +protected lemma bot_mem : ⊥ ∈ f.invtSubmodule := by simp [invtSubmodule] + +instance : BoundedOrder (f.invtSubmodule) where + top := ⟨⊤, invtSubmodule.top_mem f⟩ + bot := ⟨⊥, invtSubmodule.bot_mem f⟩ + le_top := fun ⟨p, hp⟩ ↦ by simp + bot_le := fun ⟨p, hp⟩ ↦ by simp + +@[simp] +protected lemma zero : + (0 : End R M).invtSubmodule = ⊤ := + eq_top_iff.mpr fun x ↦ by simp [invtSubmodule] + +@[simp] +protected lemma id : + invtSubmodule (LinearMap.id : End R M) = ⊤ := + eq_top_iff.mpr fun x ↦ by simp [invtSubmodule] + +protected lemma mk_eq_bot_iff {p : Submodule R M} (hp : p ∈ f.invtSubmodule) : + (⟨p, hp⟩ : f.invtSubmodule) = ⊥ ↔ p = ⊥ := + Subtype.mk_eq_bot_iff (by simp [invtSubmodule]) _ + +protected lemma mk_eq_top_iff {p : Submodule R M} (hp : p ∈ f.invtSubmodule) : + (⟨p, hp⟩ : f.invtSubmodule) = ⊤ ↔ p = ⊤ := + Subtype.mk_eq_top_iff (by simp [invtSubmodule]) _ + +@[simp] +protected lemma disjoint_mk_iff {p q : Submodule R M} + (hp : p ∈ f.invtSubmodule) (hq : q ∈ f.invtSubmodule) : + Disjoint (α := f.invtSubmodule) ⟨p, hp⟩ ⟨q, hq⟩ ↔ Disjoint p q := by + rw [disjoint_iff, disjoint_iff, Sublattice.mk_inf_mk, + Subtype.mk_eq_bot_iff (⊥ : f.invtSubmodule).property] + +protected lemma disjoint_iff {p q : f.invtSubmodule} : + Disjoint p q ↔ Disjoint (p : Submodule R M) (q : Submodule R M) := by + obtain ⟨p, hp⟩ := p + obtain ⟨q, hq⟩ := q + simp + +@[simp] +protected lemma codisjoint_mk_iff {p q : Submodule R M} + (hp : p ∈ f.invtSubmodule) (hq : q ∈ f.invtSubmodule) : + Codisjoint (α := f.invtSubmodule) ⟨p, hp⟩ ⟨q, hq⟩ ↔ Codisjoint p q := by + rw [codisjoint_iff, codisjoint_iff, Sublattice.mk_sup_mk, + Subtype.mk_eq_top_iff (⊤ : f.invtSubmodule).property] + +protected lemma codisjoint_iff {p q : f.invtSubmodule} : + Codisjoint p q ↔ Codisjoint (p : Submodule R M) (q : Submodule R M) := by + obtain ⟨p, hp⟩ := p + obtain ⟨q, hq⟩ := q + simp + +@[simp] +protected lemma isCompl_mk_iff {p q : Submodule R M} + (hp : p ∈ f.invtSubmodule) (hq : q ∈ f.invtSubmodule) : + IsCompl (α := f.invtSubmodule) ⟨p, hp⟩ ⟨q, hq⟩ ↔ IsCompl p q := by + simp [isCompl_iff] + +protected lemma isCompl_iff {p q : f.invtSubmodule} : + IsCompl p q ↔ IsCompl (p : Submodule R M) (q : Submodule R M) := by + obtain ⟨p, hp⟩ := p + obtain ⟨q, hq⟩ := q + simp + +lemma map_subtype_mem_of_mem_invtSubmodule {p : Submodule R M} (hp : p ∈ f.invtSubmodule) + {q : Submodule R p} (hq : q ∈ invtSubmodule (LinearMap.restrict f hp)) : + Submodule.map p.subtype q ∈ f.invtSubmodule := by + rintro - ⟨⟨x, hx⟩, hx', rfl⟩ + specialize hq hx' + rw [Submodule.mem_comap, LinearMap.restrict_apply] at hq + simpa [hq] using hp hx + +end invtSubmodule + +end Module.End diff --git a/Mathlib/Algebra/Module/Submodule/Map.lean b/Mathlib/Algebra/Module/Submodule/Map.lean index 84ea00362c6d2..022184894e59c 100644 --- a/Mathlib/Algebra/Module/Submodule/Map.lean +++ b/Mathlib/Algebra/Module/Submodule/Map.lean @@ -364,10 +364,16 @@ def orderIsoMapComapOfBijective [FunLike F M M₂] [SemilinearMapClass F σ₁ map_rel_iff' := map_le_map_iff_of_injective hf.injective _ _ /-- A linear isomorphism induces an order isomorphism of submodules. -/ -@[simps! symm_apply apply] +@[simps! apply] def orderIsoMapComap [EquivLike F M M₂] [SemilinearMapClass F σ₁₂ M M₂] (f : F) : Submodule R M ≃o Submodule R₂ M₂ := orderIsoMapComapOfBijective f (EquivLike.bijective f) +@[simp] +lemma orderIsoMapComap_symm_apply [EquivLike F M M₂] [SemilinearMapClass F σ₁₂ M M₂] + (f : F) (p : Submodule R₂ M₂) : + (orderIsoMapComap f).symm p = comap f p := + rfl + end OrderIso variable {F : Type*} [FunLike F M M₂] [SemilinearMapClass F σ₁₂ M M₂] diff --git a/Mathlib/Algebra/Module/Submodule/Range.lean b/Mathlib/Algebra/Module/Submodule/Range.lean index 1053c1cf8c54c..44306a514a638 100644 --- a/Mathlib/Algebra/Module/Submodule/Range.lean +++ b/Mathlib/Algebra/Module/Submodule/Range.lean @@ -277,6 +277,12 @@ theorem comap_subtype_eq_top {p p' : Submodule R M} : comap p.subtype p' = ⊤ theorem comap_subtype_self : comap p.subtype p = ⊤ := comap_subtype_eq_top.2 le_rfl +@[simp] +lemma comap_subtype_le_iff {p q r : Submodule R M} : + q.comap p.subtype ≤ r.comap p.subtype ↔ p ⊓ q ≤ p ⊓ r := + ⟨fun h ↦ by simpa using map_mono (f := p.subtype) h, + fun h ↦ by simpa using comap_mono (f := p.subtype) h⟩ + theorem range_inclusion (p q : Submodule R M) (h : p ≤ q) : range (inclusion h) = comap q.subtype p := by rw [← map_top, inclusion, LinearMap.map_codRestrict, map_top, range_subtype] diff --git a/Mathlib/Algebra/Polynomial/Module/AEval.lean b/Mathlib/Algebra/Polynomial/Module/AEval.lean index 17b934cb8564e..b9ebc29dbacfa 100644 --- a/Mathlib/Algebra/Polynomial/Module/AEval.lean +++ b/Mathlib/Algebra/Polynomial/Module/AEval.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Richard M. Hill. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Richard M. Hill -/ +import Mathlib.Algebra.Module.Submodule.Invariant import Mathlib.Algebra.Polynomial.AlgebraMap import Mathlib.RingTheory.Finiteness @@ -94,6 +95,15 @@ def _root_.LinearMap.ofAEval {N} [AddCommMonoid N] [Module R N] [Module R[X] N] LinearMap.comp_apply, LinearEquiv.coe_toLinearMap] at h ⊢ simp_rw [pow_succ, ← mul_assoc, mul_smul _ X, ← hf, ← of_symm_X_smul, ← h] +/-- Construct an `R[X]`-linear equivalence out of `AEval R M a` from a `R`-linear map out of `M`. -/ +def _root_.LinearEquiv.ofAEval {N} [AddCommMonoid N] [Module R N] [Module R[X] N] + [IsScalarTower R R[X] N] (f : M ≃ₗ[R] N) (hf : ∀ m : M, f (a • m) = (X : R[X]) • f m) : + AEval R M a ≃ₗ[R[X]] N where + __ := LinearMap.ofAEval a f hf + invFun := (of R M a) ∘ f.symm + left_inv x := by simp [LinearMap.ofAEval] + right_inv x := by simp [LinearMap.ofAEval] + lemma annihilator_eq_ker_aeval [FaithfulSMul A M] : annihilator R[X] (AEval R M a) = RingHom.ker (aeval a) := by ext p @@ -111,60 +121,51 @@ lemma annihilator_top_eq_ker_aeval [FaithfulSMul A M] : section Submodule -variable {p : Submodule R M} {q : Submodule R[X] <| AEval R M a} - -variable (R M) in -/-- We can turn an `R[X]`-submodule into an `R`-submodule by forgetting the action of `X`. -/ -def comapSubmodule : - CompleteLatticeHom (Submodule R[X] <| AEval R M a) (Submodule R M) := - (Submodule.orderIsoMapComap (of R M a)).symm.toCompleteLatticeHom.comp <| - Submodule.restrictScalarsLatticeHom R R[X] (AEval R M a) - -@[simp] lemma mem_comapSubmodule {x : M} : - x ∈ comapSubmodule R M a q ↔ of R M a x ∈ q := - Iff.rfl +variable (R M) -@[simp] lemma comapSubmodule_le_comap : - comapSubmodule R M a q ≤ (comapSubmodule R M a q).comap (Algebra.lsmul R R M a) := by - intro m hm - simpa only [Submodule.mem_comap, Algebra.lsmul_coe, mem_comapSubmodule, ← X_smul_of] using - q.smul_mem (X : R[X]) hm - -/-- An `R`-submodule which is stable under the action of `a` can be promoted to an -`R[X]`-submodule. -/ -def mapSubmodule (hp : p ≤ p.comap (Algebra.lsmul R R M a)) : Submodule R[X] <| AEval R M a := - { toAddSubmonoid := p.toAddSubmonoid.map (of R M a) - smul_mem' := by - rintro f - ⟨m : M, h : m ∈ p, rfl⟩ - simp only [AddSubsemigroup.mem_carrier, AddSubmonoid.mem_toSubsemigroup, AddSubmonoid.mem_map, - Submodule.mem_toAddSubmonoid] - exact ⟨aeval a f • m, aeval_apply_smul_mem_of_le_comap' h f a hp, of_aeval_smul a f m⟩ } - -variable (hp : p ≤ p.comap (Algebra.lsmul R R M a)) - -@[simp] lemma mem_mapSubmodule {m : AEval R M a} : - m ∈ mapSubmodule a hp ↔ (of R M a).symm m ∈ p := +/-- The natural order isomorphism between the two ways to represent invariant submodules. -/ +def mapSubmodule : + (Algebra.lsmul R R M a).invtSubmodule ≃o Submodule R[X] (AEval R M a) where + toFun p := + { toAddSubmonoid := (p : Submodule R M).toAddSubmonoid.map (of R M a) + smul_mem' := by + rintro f - ⟨m : M, h : m ∈ (p : Submodule R M), rfl⟩ + simp only [AddSubsemigroup.mem_carrier, AddSubmonoid.mem_toSubsemigroup, + AddSubmonoid.mem_map, Submodule.mem_toAddSubmonoid] + exact ⟨aeval a f • m, aeval_apply_smul_mem_of_le_comap' h f a p.2, of_aeval_smul a f m⟩ } + invFun q := ⟨(Submodule.orderIsoMapComap (of R M a)).symm (q.restrictScalars R), fun m hm ↦ by + simpa [← X_smul_of] using q.smul_mem (X : R[X]) hm⟩ + left_inv p := by ext; simp + right_inv q := by ext; aesop + map_rel_iff' {p p'} := ⟨fun h x hx ↦ by aesop, fun h x hx ↦ by aesop⟩ + +@[simp] lemma mem_mapSubmodule_apply {p : (Algebra.lsmul R R M a).invtSubmodule} {m : AEval R M a} : + m ∈ mapSubmodule R M a p ↔ (of R M a).symm m ∈ (p : Submodule R M) := ⟨fun ⟨_, hm, hm'⟩ ↦ hm'.symm ▸ hm, fun hm ↦ ⟨(of R M a).symm m, hm, rfl⟩⟩ -@[simp] lemma mapSubmodule_comapSubmodule (h := comapSubmodule_le_comap a) : - mapSubmodule a (p := comapSubmodule R M a q) h = q := by - ext; simp - -@[simp] lemma comapSubmodule_mapSubmodule : - comapSubmodule R M a (mapSubmodule a hp) = p := by - ext; simp - -variable (R M) - -lemma injective_comapSubmodule : Injective (comapSubmodule R M a) := by - intro q₁ q₂ hq - rw [← mapSubmodule_comapSubmodule (q := q₁), ← mapSubmodule_comapSubmodule (q := q₂)] - simp_rw [hq] +@[simp] lemma mem_mapSubmodule_symm_apply {q : Submodule R[X] (AEval R M a)} {m : M} : + m ∈ ((mapSubmodule R M a).symm q : Submodule R M) ↔ of R M a m ∈ q := + Iff.rfl -lemma range_comapSubmodule : - range (comapSubmodule R M a) = {p | p ≤ p.comap (Algebra.lsmul R R M a)} := - le_antisymm (fun _ ⟨_, hq⟩ ↦ hq ▸ comapSubmodule_le_comap a) - (fun _ hp ↦ ⟨mapSubmodule a hp, comapSubmodule_mapSubmodule a hp⟩) +variable {R M} +variable (p : Submodule R M) (hp : p ∈ (Algebra.lsmul R R M a).invtSubmodule) + +/-- The natural `R`-linear equivalence between the two ways to represent an invariant submodule. -/ +def equiv_mapSubmodule : + p ≃ₗ[R] mapSubmodule R M a ⟨p, hp⟩ where + toFun x := ⟨of R M a x, by simp⟩ + invFun x := ⟨((of R M _).symm (x : AEval R M a)), by obtain ⟨x, hx⟩ := x; simpa using hx⟩ + left_inv x := rfl + right_inv x := rfl + map_add' x y := rfl + map_smul' t x := rfl + +/-- The natural `R[X]`-linear equivalence between the two ways to represent an invariant submodule. +-/ +noncomputable def restrict_equiv_mapSubmodule : + (AEval R p <| (Algebra.lsmul R R M a).restrict hp) ≃ₗ[R[X]] mapSubmodule R M a ⟨p, hp⟩ := + LinearEquiv.ofAEval ((Algebra.lsmul R R M a).restrict hp) (equiv_mapSubmodule a p hp) + (fun x ↦ by simp [equiv_mapSubmodule, X_smul_of]) end Submodule diff --git a/Mathlib/LinearAlgebra/Eigenspace/Semisimple.lean b/Mathlib/LinearAlgebra/Eigenspace/Semisimple.lean index aefced273795c..96229aa54786f 100644 --- a/Mathlib/LinearAlgebra/Eigenspace/Semisimple.lean +++ b/Mathlib/LinearAlgebra/Eigenspace/Semisimple.lean @@ -25,9 +25,9 @@ namespace Module.End variable {R M : Type*} [CommRing R] [AddCommGroup M] [Module R M] {f g : End R M} -lemma apply_eq_of_mem_genEigenspace_of_comm_of_isSemisimple_of_isNilpotent_sub +lemma apply_eq_of_mem_of_comm_of_isFinitelySemisimple_of_isNil {μ : R} {k : ℕ} {m : M} (hm : m ∈ f.genEigenspace μ k) - (hfg : Commute f g) (hss : g.IsSemisimple) (hnil : IsNilpotent (f - g)) : + (hfg : Commute f g) (hss : g.IsFinitelySemisimple) (hnil : IsNilpotent (f - g)) : g m = μ • m := by set p := f.genEigenspace μ k have h₁ : MapsTo g p p := mapsTo_genEigenspace_of_comm hfg μ k @@ -41,19 +41,19 @@ lemma apply_eq_of_mem_genEigenspace_of_comm_of_isSemisimple_of_isNilpotent_sub (Commute.sub_right rfl hfg).sub_left <| Algebra.commute_algebraMap_left μ (f - g) suffices IsNilpotent ((g - algebraMap R (End R M) μ).restrict h₂) by replace this : g.restrict h₁ - algebraMap R (End R p) μ = 0 := - eq_zero_of_isNilpotent_isSemisimple this (by simpa using hss.restrict) + eq_zero_of_isNilpotent_of_isFinitelySemisimple this (by simpa using hss.restrict _) simpa [LinearMap.restrict_apply, sub_eq_zero] using LinearMap.congr_fun this ⟨m, hm⟩ simpa [LinearMap.restrict_sub h₄ h₃] using (LinearMap.restrict_commute hfg h₄ h₃).isNilpotent_sub (f.isNilpotent_restrict_sub_algebraMap μ k) (Module.End.isNilpotent.restrict h₃ hnil) -lemma IsSemisimple.genEigenspace_eq_eigenspace - (hf : f.IsSemisimple) (μ : R) {k : ℕ} (hk : 0 < k) : +lemma IsFinitelySemisimple.genEigenspace_eq_eigenspace + (hf : f.IsFinitelySemisimple) (μ : R) {k : ℕ} (hk : 0 < k) : f.genEigenspace μ k = f.eigenspace μ := by refine le_antisymm (fun m hm ↦ mem_eigenspace_iff.mpr ?_) (eigenspace_le_genEigenspace hk) - exact apply_eq_of_mem_genEigenspace_of_comm_of_isSemisimple_of_isNilpotent_sub hm rfl hf (by simp) + exact apply_eq_of_mem_of_comm_of_isFinitelySemisimple_of_isNil hm rfl hf (by simp) -lemma IsSemisimple.maxGenEigenspace_eq_eigenspace - (hf : f.IsSemisimple) (μ : R) : +lemma IsFinitelySemisimple.maxGenEigenspace_eq_eigenspace + (hf : f.IsFinitelySemisimple) (μ : R) : f.maxGenEigenspace μ = f.eigenspace μ := by simp_rw [maxGenEigenspace_def, ← (f.genEigenspace μ).monotone.iSup_nat_add 1, hf.genEigenspace_eq_eigenspace μ (Nat.zero_lt_succ _), ciSup_const] diff --git a/Mathlib/LinearAlgebra/Semisimple.lean b/Mathlib/LinearAlgebra/Semisimple.lean index 3aeb94be63ed6..3ec763f056da0 100644 --- a/Mathlib/LinearAlgebra/Semisimple.lean +++ b/Mathlib/LinearAlgebra/Semisimple.lean @@ -6,7 +6,6 @@ Authors: Oliver Nash import Mathlib.FieldTheory.Perfect import Mathlib.LinearAlgebra.Basis.VectorSpace import Mathlib.LinearAlgebra.AnnihilatingPolynomial -import Mathlib.Order.CompleteSublattice import Mathlib.RingTheory.Artinian import Mathlib.RingTheory.QuotientNilpotent import Mathlib.RingTheory.SimpleModule @@ -49,22 +48,61 @@ namespace Module.End section CommRing -variable (f g : End R M) +variable (f : End R M) /-- A linear endomorphism of an `R`-module `M` is called *semisimple* if the induced `R[X]`-module structure on `M` is semisimple. This is equivalent to saying that every `f`-invariant `R`-submodule of `M` has an `f`-invariant complement: see `Module.End.isSemisimple_iff`. -/ -abbrev IsSemisimple := IsSemisimpleModule R[X] (AEval' f) +def IsSemisimple := IsSemisimpleModule R[X] (AEval' f) -variable {f g} +/-- A weaker version of semisimplicity that only prescribes behaviour on finitely-generated +submodules. -/ +def IsFinitelySemisimple : Prop := + ∀ p (hp : p ∈ invtSubmodule f), Module.Finite R p → IsSemisimple (LinearMap.restrict f hp) + +variable {f} + +/-- A linear endomorphism is semisimple if every invariant submodule has in invariant complement. + +See also `Module.End.isSemisimple_iff`. -/ +lemma isSemisimple_iff' : + f.IsSemisimple ↔ ∀ p : invtSubmodule f, ∃ q : invtSubmodule f, IsCompl p q := by + rw [IsSemisimple, IsSemisimpleModule, (AEval.mapSubmodule R M f).symm.complementedLattice_iff, + complementedLattice_iff] + rfl lemma isSemisimple_iff : - f.IsSemisimple ↔ ∀ p : Submodule R M, p ≤ p.comap f → ∃ q, q ≤ q.comap f ∧ IsCompl p q := by - set s := (AEval.comapSubmodule R M f).range - have h : s = {p : Submodule R M | p ≤ p.comap f} := AEval.range_comapSubmodule R M f - let e := CompleteLatticeHom.toOrderIsoRangeOfInjective _ (AEval.injective_comapSubmodule R M f) - simp_rw [Module.End.IsSemisimple, IsSemisimpleModule, e.complementedLattice_iff, - s.isComplemented_iff, ← SetLike.mem_coe, h, mem_setOf_eq] + f.IsSemisimple ↔ ∀ p ∈ invtSubmodule f, ∃ q ∈ invtSubmodule f, IsCompl p q := by + simp_rw [isSemisimple_iff'] + aesop + +lemma isSemisimple_restrict_iff (p) (hp : p ∈ invtSubmodule f) : + IsSemisimple (LinearMap.restrict f hp) ↔ + ∀ q ∈ f.invtSubmodule, q ≤ p → ∃ r ≤ p, r ∈ f.invtSubmodule ∧ Disjoint q r ∧ q ⊔ r = p := by + let e : Submodule R[X] (AEval' (f.restrict hp)) ≃o Iic (AEval.mapSubmodule R M f ⟨p, hp⟩) := + (Submodule.orderIsoMapComap <| AEval.restrict_equiv_mapSubmodule f p hp).trans + (Submodule.mapIic _) + simp_rw [IsSemisimple, IsSemisimpleModule, e.complementedLattice_iff, disjoint_iff, + ← (OrderIso.Iic _ _).complementedLattice_iff, Iic.complementedLattice_iff, Subtype.forall, + Subtype.exists, Subtype.mk_le_mk, Sublattice.mk_inf_mk, Sublattice.mk_sup_mk, Subtype.mk.injEq, + exists_and_left, exists_and_right, invtSubmodule.mk_eq_bot_iff, exists_prop, and_assoc] + rfl + +/-- A linear endomorphism is finitely semisimple if it is semisimple on every finitely-generated +invariant submodule. + +See also `Module.End.isFinitelySemisimple_iff`. -/ +lemma isFinitelySemisimple_iff' : + f.IsFinitelySemisimple ↔ ∀ p (hp : p ∈ invtSubmodule f), + Module.Finite R p → IsSemisimple (LinearMap.restrict f hp) := + Iff.rfl + +/-- A characterisation of `Module.End.IsFinitelySemisimple` using only the lattice of submodules of +`M` (thus avoiding submodules of submodules). -/ +lemma isFinitelySemisimple_iff : + f.IsFinitelySemisimple ↔ ∀ p ∈ invtSubmodule f, Module.Finite R p → ∀ q ∈ invtSubmodule f, + q ≤ p → ∃ r, r ≤ p ∧ r ∈ invtSubmodule f ∧ Disjoint q r ∧ q ⊔ r = p := by + simp_rw [isFinitelySemisimple_iff', isSemisimple_restrict_iff] @[simp] lemma isSemisimple_zero [IsSemisimpleModule R M] : IsSemisimple (0 : Module.End R M) := by @@ -74,7 +112,16 @@ lemma isSemisimple_zero [IsSemisimpleModule R M] : IsSemisimple (0 : Module.End lemma isSemisimple_id [IsSemisimpleModule R M] : IsSemisimple (LinearMap.id : Module.End R M) := by simpa [isSemisimple_iff] using exists_isCompl -@[simp] lemma isSemisimple_neg : (-f).IsSemisimple ↔ f.IsSemisimple := by simp [isSemisimple_iff] +@[simp] lemma isSemisimple_neg : (-f).IsSemisimple ↔ f.IsSemisimple := by + simp [isSemisimple_iff, mem_invtSubmodule] + +variable (f) in +protected lemma _root_.LinearEquiv.isSemisimple_iff {M₂ : Type*} [AddCommGroup M₂] [Module R M₂] + (g : End R M₂) (e : M ≃ₗ[R] M₂) (he : e ∘ₗ f = g ∘ₗ e) : + f.IsSemisimple ↔ g.IsSemisimple := by + let e : AEval' f ≃ₗ[R[X]] AEval' g := LinearEquiv.ofAEval _ (e.trans (AEval'.of g)) fun x ↦ by + simpa [AEval'.X_smul_of] using LinearMap.congr_fun he x + exact (Submodule.orderIsoMapComap e).complementedLattice_iff lemma eq_zero_of_isNilpotent_isSemisimple (hn : IsNilpotent f) (hs : f.IsSemisimple) : f = 0 := by have ⟨n, h0⟩ := hn @@ -82,24 +129,74 @@ lemma eq_zero_of_isNilpotent_isSemisimple (hn : IsNilpotent f) (hs : f.IsSemisim rw [← RingHom.mem_ker, ← AEval.annihilator_eq_ker_aeval (M := M)] at h0 ⊢ exact hs.annihilator_isRadical _ _ ⟨n, h0⟩ +lemma eq_zero_of_isNilpotent_of_isFinitelySemisimple + (hn : IsNilpotent f) (hs : IsFinitelySemisimple f) : f = 0 := by + have (p) (hp₁ : p ∈ f.invtSubmodule) (hp₂ : Module.Finite R p) : f.restrict hp₁ = 0 := by + specialize hs p hp₁ hp₂ + replace hn : IsNilpotent (f.restrict hp₁) := isNilpotent.restrict hp₁ hn + exact eq_zero_of_isNilpotent_isSemisimple hn hs + ext x + obtain ⟨k : ℕ, hk : f ^ k = 0⟩ := hn + let p := Submodule.span R {(f ^ i) x | (i : ℕ) (_ : i ≤ k)} + have hp₁ : p ∈ f.invtSubmodule := by + simp only [mem_invtSubmodule, p, Submodule.span_le] + rintro - ⟨i, hi, rfl⟩ + apply Submodule.subset_span + rcases lt_or_eq_of_le hi with hik | rfl + · exact ⟨i + 1, hik, by simpa [LinearMap.pow_apply] using iterate_succ_apply' f i x⟩ + · exact ⟨i, by simp [hk]⟩ + have hp₂ : Module.Finite R p := by + let g : ℕ → M := fun i ↦ (f ^ i) x + have hg : {(f ^ i) x | (i : ℕ) (_ : i ≤ k)} = g '' Iic k := by ext; simp [g] + exact Module.Finite.span_of_finite _ <| hg ▸ toFinite (g '' Iic k) + simpa [LinearMap.restrict_apply, Subtype.ext_iff] using + LinearMap.congr_fun (this p hp₁ hp₂) ⟨x, Submodule.subset_span ⟨0, k.zero_le, rfl⟩⟩ + @[simp] lemma isSemisimple_sub_algebraMap_iff {μ : R} : (f - algebraMap R (End R M) μ).IsSemisimple ↔ f.IsSemisimple := by suffices ∀ p : Submodule R M, p ≤ p.comap (f - algebraMap R (Module.End R M) μ) ↔ p ≤ p.comap f by - simp [isSemisimple_iff, this] + simp [mem_invtSubmodule, isSemisimple_iff, this] refine fun p ↦ ⟨fun h x hx ↦ ?_, fun h x hx ↦ p.sub_mem (h hx) (p.smul_mem μ hx)⟩ simpa using p.add_mem (h hx) (p.smul_mem μ hx) -lemma IsSemisimple.restrict {p : Submodule R M} {hp : MapsTo f p p} (hf : f.IsSemisimple) : +lemma IsSemisimple.restrict {p : Submodule R M} (hp : p ∈ f.invtSubmodule) (hf : f.IsSemisimple) : IsSemisimple (f.restrict hp) := by - simp only [isSemisimple_iff] at hf ⊢ - intro q hq - replace hq : MapsTo f (q.map p.subtype) (q.map p.subtype) := by - rintro - ⟨⟨x, hx⟩, hx', rfl⟩; exact ⟨⟨f x, hp hx⟩, by simpa using hq hx', rfl⟩ - obtain ⟨r, hr₁, hr₂⟩ := hf _ hq - refine ⟨r.comap p.subtype, fun x hx ↦ hr₁ hx, ?_⟩ - rw [← q.comap_map_eq_of_injective p.injective_subtype] - exact p.isCompl_comap_subtype_of_isCompl_of_le hr₂ <| p.map_subtype_le q + rw [IsSemisimple] at hf ⊢ + let e : Submodule R[X] (AEval' (LinearMap.restrict f hp)) ≃o + Iic (AEval.mapSubmodule R M f ⟨p, hp⟩) := + (Submodule.orderIsoMapComap <| AEval.restrict_equiv_mapSubmodule f p hp).trans <| + Submodule.mapIic _ + exact e.complementedLattice_iff.mpr inferInstance + +lemma IsSemisimple.isFinitelySemisimple (hf : f.IsSemisimple) : + f.IsFinitelySemisimple := + isFinitelySemisimple_iff'.mp fun _ _ _ ↦ hf.restrict _ + +@[simp] +lemma isFinitelySemisimple_iff_isSemisimple [Module.Finite R M] : + f.IsFinitelySemisimple ↔ f.IsSemisimple := by + refine ⟨fun hf ↦ isSemisimple_iff.mpr fun p hp ↦ ?_, IsSemisimple.isFinitelySemisimple⟩ + obtain ⟨q, -, hq₁, hq₂, hq₃⟩ := + isFinitelySemisimple_iff.mp hf ⊤ (invtSubmodule.top_mem f) inferInstance p hp le_top + exact ⟨q, hq₁, hq₂, codisjoint_iff.mpr hq₃⟩ + +@[simp] +lemma isFinitelySemisimple_sub_algebraMap_iff {μ : R} : + (f - algebraMap R (End R M) μ).IsFinitelySemisimple ↔ f.IsFinitelySemisimple := by + suffices ∀ p : Submodule R M, p ≤ p.comap (f - algebraMap R (Module.End R M) μ) ↔ p ≤ p.comap f by + simp_rw [isFinitelySemisimple_iff, mem_invtSubmodule, this] + refine fun p ↦ ⟨fun h x hx ↦ ?_, fun h x hx ↦ p.sub_mem (h hx) (p.smul_mem μ hx)⟩ + simpa using p.add_mem (h hx) (p.smul_mem μ hx) + +lemma IsFinitelySemisimple.restrict {p : Submodule R M} (hp : p ∈ f.invtSubmodule) + (hf : f.IsFinitelySemisimple) : + IsFinitelySemisimple (f.restrict hp) := by + intro q hq₁ hq₂ + have := invtSubmodule.map_subtype_mem_of_mem_invtSubmodule f hp hq₁ + let e : q ≃ₗ[R] q.map p.subtype := p.equivSubtypeMap q + rw [e.isSemisimple_iff ((LinearMap.restrict f hp).restrict hq₁) (LinearMap.restrict f this) rfl] + exact hf _ this (Finite.map q p.subtype) end CommRing @@ -109,7 +206,7 @@ variable {K : Type*} [Field K] [Module K M] {f g : End K M} lemma IsSemisimple_smul_iff {t : K} (ht : t ≠ 0) : (t • f).IsSemisimple ↔ f.IsSemisimple := by - simp [isSemisimple_iff, Submodule.comap_smul f (h := ht)] + simp [isSemisimple_iff, mem_invtSubmodule, Submodule.comap_smul f (h := ht)] lemma IsSemisimple_smul (t : K) (h : f.IsSemisimple) : (t • f).IsSemisimple := by diff --git a/Mathlib/Order/Disjoint.lean b/Mathlib/Order/Disjoint.lean index 959e072e1b1b6..b49a00f24c125 100644 --- a/Mathlib/Order/Disjoint.lean +++ b/Mathlib/Order/Disjoint.lean @@ -586,6 +586,10 @@ class ComplementedLattice (α) [Lattice α] [BoundedOrder α] : Prop where /-- In a `ComplementedLattice`, every element admits a complement. -/ exists_isCompl : ∀ a : α, ∃ b : α, IsCompl a b +lemma complementedLattice_iff (α) [Lattice α] [BoundedOrder α] : + ComplementedLattice α ↔ ∀ a : α, ∃ b : α, IsCompl a b := + ⟨fun ⟨h⟩ ↦ h, fun h ↦ ⟨h⟩⟩ + export ComplementedLattice (exists_isCompl) instance Subsingleton.instComplementedLattice diff --git a/Mathlib/Order/Hom/Set.lean b/Mathlib/Order/Hom/Set.lean index 382a2354e4f5d..df6730da16330 100644 --- a/Mathlib/Order/Hom/Set.lean +++ b/Mathlib/Order/Hom/Set.lean @@ -8,13 +8,14 @@ import Mathlib.Logic.Equiv.Set import Mathlib.Data.Set.Monotone import Mathlib.Data.Set.Image import Mathlib.Order.WellFounded +import Mathlib.Order.Interval.Set.Basic /-! # Order homomorphisms and sets -/ -open OrderDual +open OrderDual Set variable {α β : Type*} @@ -159,6 +160,36 @@ instance subsingleton_of_wellFoundedGT' [LinearOrder β] [WellFoundedGT β] [Pre instance unique_of_wellFoundedGT [LinearOrder α] [WellFoundedGT α] : Unique (α ≃o α) := Unique.mk' _ +/-- An order isomorphism between lattices induces an order isomorphism between corresponding +interval sublattices. -/ +protected def Iic [Lattice α] [Lattice β] (e : α ≃o β) (x : α) : + Iic x ≃o Iic (e x) where + toFun y := ⟨e y, (map_le_map_iff _).mpr y.property⟩ + invFun y := ⟨e.symm y, e.symm_apply_le.mpr y.property⟩ + left_inv y := by simp + right_inv y := by simp + map_rel_iff' := by simp + +/-- An order isomorphism between lattices induces an order isomorphism between corresponding +interval sublattices. -/ +protected def Ici [Lattice α] [Lattice β] (e : α ≃o β) (x : α) : + Ici x ≃o Ici (e x) where + toFun y := ⟨e y, (map_le_map_iff _).mpr y.property⟩ + invFun y := ⟨e.symm y, e.le_symm_apply.mpr y.property⟩ + left_inv y := by simp + right_inv y := by simp + map_rel_iff' := by simp + +/-- An order isomorphism between lattices induces an order isomorphism between corresponding +interval sublattices. -/ +protected def Icc [Lattice α] [Lattice β] (e : α ≃o β) (x y : α) : + Icc x y ≃o Icc (e x) (e y) where + toFun z := ⟨e z, by simp only [mem_Icc, map_le_map_iff]; exact z.property⟩ + invFun z := ⟨e.symm z, by simp only [mem_Icc, e.le_symm_apply, e.symm_apply_le]; exact z.property⟩ + left_inv y := by simp + right_inv y := by simp + map_rel_iff' := by simp + end OrderIso section BooleanAlgebra diff --git a/Mathlib/Order/LatticeIntervals.lean b/Mathlib/Order/LatticeIntervals.lean index 7be5390c9e2a4..92761688d3430 100644 --- a/Mathlib/Order/LatticeIntervals.lean +++ b/Mathlib/Order/LatticeIntervals.lean @@ -126,6 +126,16 @@ protected lemma isCompl_iff [Lattice α] [BoundedOrder α] {x y : Iic a} : IsCompl x y ↔ Disjoint (x : α) (y : α) ∧ (x : α) ⊔ (y : α) = a := by rw [_root_.isCompl_iff, Iic.disjoint_iff, Iic.codisjoint_iff] +protected lemma complementedLattice_iff [Lattice α] [BoundedOrder α] : + ComplementedLattice (Iic a) ↔ ∀ b, b ≤ a → ∃ c ≤ a, b ⊓ c = ⊥ ∧ b ⊔ c = a := by + refine ⟨fun h b hb ↦ ?_, fun h ↦ ⟨fun ⟨x, hx⟩ ↦ ?_⟩⟩ + · obtain ⟨⟨c, hc₁⟩, hc⟩ := exists_isCompl (⟨b, hb⟩ : Iic a) + obtain ⟨hc₂, hc₃⟩ := Set.Iic.isCompl_iff.mp hc + exact ⟨c, hc₁, disjoint_iff.mp hc₂, hc₃⟩ + · simp_rw [Set.Iic.isCompl_iff] + obtain ⟨c, hc₁, hc₂, hc₃⟩ := h x hx + exact ⟨⟨c, hc₁⟩, disjoint_iff.mpr hc₂, hc₃⟩ + end Iic namespace Ici From 616d464002748587cb49debd1bac547cb87f3d84 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 22 Oct 2024 08:37:14 +0000 Subject: [PATCH 259/505] chore(BorelSpace/Order): allow anonymous dot notation on `Measurable` lemma (#17814) From GibbsMeasure --- .../Constructions/BorelSpace/Order.lean | 132 ++++++++++++------ .../Constructions/BorelSpace/Real.lean | 8 +- .../Constructions/Prod/Basic.lean | 3 +- .../MeasureTheory/Decomposition/Lebesgue.lean | 4 +- .../Function/AEMeasurableOrder.lean | 4 +- .../Function/LpSeminorm/Basic.lean | 2 +- Mathlib/MeasureTheory/Integral/Lebesgue.lean | 10 +- .../MeasureTheory/MeasurableSpace/Basic.lean | 2 +- Mathlib/MeasureTheory/Measure/GiryMonad.lean | 2 +- .../Kernel/Disintegration/CDFToKernel.lean | 2 +- .../Kernel/Disintegration/Density.lean | 2 +- .../Disintegration/MeasurableStieltjes.lean | 4 +- .../Kernel/MeasurableIntegral.lean | 2 +- .../Probability/Martingale/Upcrossing.lean | 2 +- 14 files changed, 111 insertions(+), 68 deletions(-) diff --git a/Mathlib/MeasureTheory/Constructions/BorelSpace/Order.lean b/Mathlib/MeasureTheory/Constructions/BorelSpace/Order.lean index 8a3ca4d20c7d5..c52471a68104f 100644 --- a/Mathlib/MeasureTheory/Constructions/BorelSpace/Order.lean +++ b/Mathlib/MeasureTheory/Constructions/BorelSpace/Order.lean @@ -27,10 +27,10 @@ import Mathlib.MeasureTheory.Constructions.BorelSpace.Basic intervals of the given kind. * `UpperSemicontinuous.measurable`, `LowerSemicontinuous.measurable`: Semicontinuous functions are measurable. -* `measurable_iSup`, `measurable_iInf`, `measurable_sSup`, `measurable_sInf`: +* `Measurable.iSup`, `Measurable.iInf`, `Measurable.sSup`, `Measurable.sInf`: Countable supremums and infimums of measurable functions to conditionally complete linear orders are measurable. -* `measurable_liminf`, `measurable_limsup`: +* `Measurable.liminf`, `Measurable.limsup`: Countable liminfs and limsups of measurable functions to conditionally complete linear orders are measurable. @@ -744,8 +744,8 @@ section ConditionallyCompleteLinearOrder variable [ConditionallyCompleteLinearOrder α] [OrderTopology α] [SecondCountableTopology α] -@[measurability] -theorem measurable_iSup {ι} [Countable ι] {f : ι → δ → α} (hf : ∀ i, Measurable (f i)) : +@[measurability, fun_prop] +protected theorem Measurable.iSup {ι} [Countable ι] {f : ι → δ → α} (hf : ∀ i, Measurable (f i)) : Measurable (fun b ↦ ⨆ i, f i b) := by rcases isEmpty_or_nonempty ι with hι|hι · simp [iSup_of_empty'] @@ -762,74 +762,111 @@ theorem measurable_iSup {ι} [Countable ι] {f : ι → δ → α} (hf : ∀ i, apply csSup_of_not_bddAbove exact hb +@[deprecated (since := "2024-10-21")] +alias measurable_iSup := Measurable.iSup + +-- TODO: Why does this error? +-- /-- Compositional version of `Measurable.iSup` for use by `fun_prop`. -/ +-- @[fun_prop] +-- protected lemma Measurable.iSup'' {_ : MeasurableSpace γ} {ι : Sort*} [Countable ι] +-- {f : ι → γ → δ → α} {h : γ → δ} (hf : ∀ i, Measurable ↿(f i)) (hh : Measurable h) : +-- Measurable fun a ↦ (⨆ i, f i a) (h a) := by +-- simp_rw [iSup_apply] +-- exact .iSup fun i ↦ by fun_prop + @[measurability] -theorem aemeasurable_iSup {ι} {μ : Measure δ} [Countable ι] {f : ι → δ → α} +protected theorem AEMeasurable.iSup {ι} {μ : Measure δ} [Countable ι] {f : ι → δ → α} (hf : ∀ i, AEMeasurable (f i) μ) : AEMeasurable (fun b => ⨆ i, f i b) μ := by - refine ⟨fun b ↦ ⨆ i, (hf i).mk (f i) b, measurable_iSup (fun i ↦ (hf i).measurable_mk), ?_⟩ + refine ⟨fun b ↦ ⨆ i, (hf i).mk (f i) b, .iSup (fun i ↦ (hf i).measurable_mk), ?_⟩ filter_upwards [ae_all_iff.2 (fun i ↦ (hf i).ae_eq_mk)] with b hb using by simp [hb] -@[measurability] -theorem measurable_iInf {ι} [Countable ι] {f : ι → δ → α} (hf : ∀ i, Measurable (f i)) : +@[deprecated (since := "2024-10-21")] +alias aemeasurable_iSup := AEMeasurable.iSup + +@[measurability, fun_prop] +protected theorem Measurable.iInf {ι} [Countable ι] {f : ι → δ → α} (hf : ∀ i, Measurable (f i)) : Measurable fun b => ⨅ i, f i b := - measurable_iSup (α := αᵒᵈ) hf + .iSup (α := αᵒᵈ) hf + +@[deprecated (since := "2024-10-21")] +alias measurable_iInf := Measurable.iInf @[measurability] -theorem aemeasurable_iInf {ι} {μ : Measure δ} [Countable ι] {f : ι → δ → α} +protected theorem AEMeasurable.iInf {ι} {μ : Measure δ} [Countable ι] {f : ι → δ → α} (hf : ∀ i, AEMeasurable (f i) μ) : AEMeasurable (fun b => ⨅ i, f i b) μ := - aemeasurable_iSup (α := αᵒᵈ) hf + .iSup (α := αᵒᵈ) hf + +@[deprecated (since := "2024-10-21")] +alias aemeasurable_iInf := AEMeasurable.iInf -theorem measurable_sSup {ι} {f : ι → δ → α} {s : Set ι} (hs : s.Countable) +protected theorem Measurable.sSup {ι} {f : ι → δ → α} {s : Set ι} (hs : s.Countable) (hf : ∀ i ∈ s, Measurable (f i)) : Measurable fun x => sSup ((fun i => f i x) '' s) := by - have : Countable ↑s := countable_coe_iff.2 hs - convert measurable_iSup (f := (fun (i : s) ↦ f i)) (fun i ↦ hf i i.2) using 1 - ext b - congr - exact image_eq_range (fun i ↦ f i b) s + simp_rw [image_eq_range] + have : Countable s := hs.to_subtype + exact .iSup fun i ↦ hf i i.2 -theorem measurable_sInf {ι} {f : ι → δ → α} {s : Set ι} (hs : s.Countable) +@[deprecated (since := "2024-10-21")] +alias measurable_sSup := Measurable.sSup + +protected theorem Measurable.sInf {ι} {f : ι → δ → α} {s : Set ι} (hs : s.Countable) (hf : ∀ i ∈ s, Measurable (f i)) : Measurable fun x => sInf ((fun i => f i x) '' s) := - measurable_sSup (α := αᵒᵈ) hs hf + .sSup (α := αᵒᵈ) hs hf + +@[deprecated (since := "2024-10-21")] +alias measurable_sInf := Measurable.sInf -theorem measurable_biSup {ι} (s : Set ι) {f : ι → δ → α} (hs : s.Countable) +theorem Measurable.biSup {ι} (s : Set ι) {f : ι → δ → α} (hs : s.Countable) (hf : ∀ i ∈ s, Measurable (f i)) : Measurable fun b => ⨆ i ∈ s, f i b := by haveI : Encodable s := hs.toEncodable by_cases H : ∀ i, i ∈ s · have : ∀ b, ⨆ i ∈ s, f i b = ⨆ (i : s), f i b := fun b ↦ cbiSup_eq_of_forall (f := fun i ↦ f i b) H simp only [this] - exact measurable_iSup (fun (i : s) ↦ hf i i.2) + exact .iSup (fun (i : s) ↦ hf i i.2) · have : ∀ b, ⨆ i ∈ s, f i b = (⨆ (i : s), f i b) ⊔ sSup ∅ := fun b ↦ cbiSup_eq_of_not_forall (f := fun i ↦ f i b) H simp only [this] apply Measurable.sup _ measurable_const - exact measurable_iSup (fun (i : s) ↦ hf i i.2) + exact .iSup (fun (i : s) ↦ hf i i.2) -theorem aemeasurable_biSup {ι} {μ : Measure δ} (s : Set ι) {f : ι → δ → α} (hs : s.Countable) +@[deprecated (since := "2024-10-21")] +alias measurable_biSup := Measurable.biSup + +theorem AEMeasurable.biSup {ι} {μ : Measure δ} (s : Set ι) {f : ι → δ → α} (hs : s.Countable) (hf : ∀ i ∈ s, AEMeasurable (f i) μ) : AEMeasurable (fun b => ⨆ i ∈ s, f i b) μ := by classical let g : ι → δ → α := fun i ↦ if hi : i ∈ s then (hf i hi).mk (f i) else fun _b ↦ sSup ∅ have : ∀ i ∈ s, Measurable (g i) := by intro i hi simpa [g, hi] using (hf i hi).measurable_mk - refine ⟨fun b ↦ ⨆ (i) (_ : i ∈ s), g i b, measurable_biSup s hs this, ?_⟩ + refine ⟨fun b ↦ ⨆ (i) (_ : i ∈ s), g i b, .biSup s hs this, ?_⟩ have : ∀ i ∈ s, ∀ᵐ b ∂μ, f i b = g i b := fun i hi ↦ by simpa [g, hi] using (hf i hi).ae_eq_mk filter_upwards [(ae_ball_iff hs).2 this] with b hb exact iSup_congr fun i => iSup_congr (hb i) -theorem measurable_biInf {ι} (s : Set ι) {f : ι → δ → α} (hs : s.Countable) +@[deprecated (since := "2024-10-21")] +alias aemeasurable_biSup := AEMeasurable.biSup + +theorem Measurable.biInf {ι} (s : Set ι) {f : ι → δ → α} (hs : s.Countable) (hf : ∀ i ∈ s, Measurable (f i)) : Measurable fun b => ⨅ i ∈ s, f i b := - measurable_biSup (α := αᵒᵈ) s hs hf + .biSup (α := αᵒᵈ) s hs hf -theorem aemeasurable_biInf {ι} {μ : Measure δ} (s : Set ι) {f : ι → δ → α} (hs : s.Countable) +@[deprecated (since := "2024-10-21")] +alias measurable_biInf := Measurable.biInf + +theorem AEMeasurable.biInf {ι} {μ : Measure δ} (s : Set ι) {f : ι → δ → α} (hs : s.Countable) (hf : ∀ i ∈ s, AEMeasurable (f i) μ) : AEMeasurable (fun b => ⨅ i ∈ s, f i b) μ := - aemeasurable_biSup (α := αᵒᵈ) s hs hf + .biSup (α := αᵒᵈ) s hs hf + +@[deprecated (since := "2024-10-21")] +alias aemeasurable_biInf := AEMeasurable.biInf -/-- `liminf` over a general filter is measurable. See `measurable_liminf` for the version over `ℕ`. +/-- `liminf` over a general filter is measurable. See `Measurable.liminf` for the version over `ℕ`. -/ -theorem measurable_liminf' {ι ι'} {f : ι → δ → α} {v : Filter ι} (hf : ∀ i, Measurable (f i)) +theorem Measurable.liminf' {ι ι'} {f : ι → δ → α} {v : Filter ι} (hf : ∀ i, Measurable (f i)) {p : ι' → Prop} {s : ι' → Set ι} (hv : v.HasCountableBasis p s) (hs : ∀ j, (s j).Countable) : Measurable fun x => liminf (f · x) v := by classical @@ -852,11 +889,10 @@ theorem measurable_liminf' {ι ι'} {f : ι → δ → α} {v : Filter ι} (hf : have mc_meas : MeasurableSet {x | ∀ (j : Subtype p), x ∉ m j} := by rw [setOf_forall] exact MeasurableSet.iInter (fun j ↦ (m_meas j).compl) - apply Measurable.piecewise mc_meas measurable_const - apply measurable_iSup (fun j ↦ ?_) + refine measurable_const.piecewise mc_meas <| .iSup fun j ↦ ?_ let reparam : δ → Subtype p → Subtype p := fun x ↦ liminf_reparam (fun i ↦ f i x) s p let F0 : Subtype p → δ → α := fun j x ↦ ⨅ (i : s j), f i x - have F0_meas : ∀ j, Measurable (F0 j) := fun j ↦ measurable_iInf (fun (i : s j) ↦ hf i) + have F0_meas : ∀ j, Measurable (F0 j) := fun j ↦ .iInf (fun (i : s j) ↦ hf i) set F1 : δ → α := fun x ↦ F0 (reparam x j) x with hF1 change Measurable F1 let g : ℕ → Subtype p := Classical.choose (exists_surjective_nat (Subtype p)) @@ -881,26 +917,38 @@ theorem measurable_liminf' {ι ι'} {f : ι → δ → α} {v : Filter ι} (hf : apply Measurable.find (fun n ↦ F0_meas (g n)) (fun n ↦ ?_) exact (m_meas (g n)).union mc_meas -/-- `limsup` over a general filter is measurable. See `measurable_limsup` for the version over `ℕ`. +@[deprecated (since := "2024-10-21")] +alias measurable_liminf' := Measurable.liminf' + +/-- `limsup` over a general filter is measurable. See `Measurable.limsup` for the version over `ℕ`. -/ -theorem measurable_limsup' {ι ι'} {f : ι → δ → α} {u : Filter ι} (hf : ∀ i, Measurable (f i)) +theorem Measurable.limsup' {ι ι'} {f : ι → δ → α} {u : Filter ι} (hf : ∀ i, Measurable (f i)) {p : ι' → Prop} {s : ι' → Set ι} (hu : u.HasCountableBasis p s) (hs : ∀ i, (s i).Countable) : Measurable fun x => limsup (fun i => f i x) u := - measurable_liminf' (α := αᵒᵈ) hf hu hs + .liminf' (α := αᵒᵈ) hf hu hs + +@[deprecated (since := "2024-10-21")] +alias measurable_limsup' := Measurable.limsup' -/-- `liminf` over `ℕ` is measurable. See `measurable_liminf'` for a version with a general filter. +/-- `liminf` over `ℕ` is measurable. See `Measurable.liminf'` for a version with a general filter. -/ @[measurability] -theorem measurable_liminf {f : ℕ → δ → α} (hf : ∀ i, Measurable (f i)) : +theorem Measurable.liminf {f : ℕ → δ → α} (hf : ∀ i, Measurable (f i)) : Measurable fun x => liminf (fun i => f i x) atTop := - measurable_liminf' hf atTop_countable_basis fun _ => to_countable _ + .liminf' hf atTop_countable_basis fun _ => to_countable _ -/-- `limsup` over `ℕ` is measurable. See `measurable_limsup'` for a version with a general filter. +@[deprecated (since := "2024-10-21")] +alias measurable_liminf := Measurable.liminf + +/-- `limsup` over `ℕ` is measurable. See `Measurable.limsup'` for a version with a general filter. -/ @[measurability] -theorem measurable_limsup {f : ℕ → δ → α} (hf : ∀ i, Measurable (f i)) : +theorem Measurable.limsup {f : ℕ → δ → α} (hf : ∀ i, Measurable (f i)) : Measurable fun x => limsup (fun i => f i x) atTop := - measurable_limsup' hf atTop_countable_basis fun _ => to_countable _ + .limsup' hf atTop_countable_basis fun _ => to_countable _ + +@[deprecated (since := "2024-10-21")] +alias measurable_limsup := Measurable.limsup end ConditionallyCompleteLinearOrder diff --git a/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean b/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean index ed5e038f6f5c4..c89fefe9a34d6 100644 --- a/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean +++ b/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean @@ -280,7 +280,7 @@ theorem measurable_of_tendsto' {ι : Type*} {f : ι → α → ℝ≥0∞} {g : exact ((lim y).comp hx).liminf_eq rw [← this] show Measurable fun y => liminf (fun n => (f (x n) y : ℝ≥0∞)) atTop - exact measurable_liminf fun n => hf (x n) + exact .liminf fun n => hf (x n) @[deprecated (since := "2024-03-09")] alias _root_.measurable_of_tendsto_ennreal' := ENNReal.measurable_of_tendsto' @@ -360,8 +360,7 @@ theorem AEMeasurable.ennreal_toReal {f : α → ℝ≥0∞} {μ : Measure α} (h theorem Measurable.ennreal_tsum {ι} [Countable ι] {f : ι → α → ℝ≥0∞} (h : ∀ i, Measurable (f i)) : Measurable fun x => ∑' i, f i x := by simp_rw [ENNReal.tsum_eq_iSup_sum] - apply measurable_iSup - exact fun s => s.measurable_sum fun i _ => h i + exact .iSup fun s ↦ s.measurable_sum fun i _ => h i @[measurability, fun_prop] theorem Measurable.ennreal_tsum' {ι} [Countable ι] {f : ι → α → ℝ≥0∞} (h : ∀ i, Measurable (f i)) : @@ -379,8 +378,7 @@ theorem Measurable.nnreal_tsum {ι} [Countable ι] {f : ι → α → ℝ≥0} ( theorem AEMeasurable.ennreal_tsum {ι} [Countable ι] {f : ι → α → ℝ≥0∞} {μ : Measure α} (h : ∀ i, AEMeasurable (f i) μ) : AEMeasurable (fun x => ∑' i, f i x) μ := by simp_rw [ENNReal.tsum_eq_iSup_sum] - apply aemeasurable_iSup - exact fun s => Finset.aemeasurable_sum s fun i _ => h i + exact .iSup fun s ↦ Finset.aemeasurable_sum s fun i _ => h i @[measurability, fun_prop] theorem AEMeasurable.nnreal_tsum {α : Type*} {_ : MeasurableSpace α} {ι : Type*} [Countable ι] diff --git a/Mathlib/MeasureTheory/Constructions/Prod/Basic.lean b/Mathlib/MeasureTheory/Constructions/Prod/Basic.lean index a41e36a07b43e..7b687ed495686 100644 --- a/Mathlib/MeasureTheory/Constructions/Prod/Basic.lean +++ b/Mathlib/MeasureTheory/Constructions/Prod/Basic.lean @@ -267,10 +267,9 @@ theorem Measurable.lintegral_prod_right' [SFinite ν] : conv => enter [1, x]; erw [lintegral_add_left (hf.comp m)] exact h2f.add h2g · intro f hf h2f h3f - have := measurable_iSup h3f have : ∀ x, Monotone fun n y => f n (x, y) := fun x i j hij y => h2f hij (x, y) conv => enter [1, x]; erw [lintegral_iSup (fun n => (hf n).comp m) (this x)] - assumption + exact .iSup h3f /-- The Lebesgue integral is measurable. This shows that the integrand of (the right-hand-side of) Tonelli's theorem is measurable. diff --git a/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean b/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean index de3c05de65b39..71f36a0f282cf 100644 --- a/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean +++ b/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean @@ -777,7 +777,7 @@ theorem iSup_mem_measurableLE (f : ℕ → α → ℝ≥0∞) (hf : ∀ n, f n (fun a : α ↦ ⨆ (k : ℕ) (_ : k ≤ m + 1), f k a) = fun a ↦ f m.succ a ⊔ ⨆ (k : ℕ) (_ : k ≤ m), f k a := funext fun _ ↦ iSup_succ_eq_sup _ _ _ - refine ⟨measurable_iSup fun n ↦ Measurable.iSup_Prop _ (hf n).1, fun A hA ↦ ?_⟩ + refine ⟨.iSup fun n ↦ Measurable.iSup_Prop _ (hf n).1, fun A hA ↦ ?_⟩ rw [this]; exact (sup_mem_measurableLE (hf m.succ) hm).2 A hA theorem iSup_mem_measurableLE' (f : ℕ → α → ℝ≥0∞) (hf : ∀ n, f n ∈ measurableLE μ ν) (n : ℕ) : @@ -846,7 +846,7 @@ theorem haveLebesgueDecomposition_of_finiteMeasure [IsFiniteMeasure μ] [IsFinit · refine Filter.Eventually.of_forall fun a ↦ ?_ simp [tendsto_atTop_iSup (iSup_monotone' f a)] have hξm : Measurable ξ := by - convert measurable_iSup fun n ↦ (iSup_mem_measurableLE _ hf₁ n).1 + convert Measurable.iSup fun n ↦ (iSup_mem_measurableLE _ hf₁ n).1 simp [hξ] -- `ξ` is the `f` in the theorem statement and we set `μ₁` to be `μ - ν.withDensity ξ` -- since we need `μ₁ + ν.withDensity ξ = μ` diff --git a/Mathlib/MeasureTheory/Function/AEMeasurableOrder.lean b/Mathlib/MeasureTheory/Function/AEMeasurableOrder.lean index 7d3d2817a1aba..0f4dba7c8de33 100644 --- a/Mathlib/MeasureTheory/Function/AEMeasurableOrder.lean +++ b/Mathlib/MeasureTheory/Function/AEMeasurableOrder.lean @@ -54,9 +54,7 @@ theorem MeasureTheory.aemeasurable_of_exist_almost_disjoint_supersets {α : Type intro i exact MeasurableSet.biInter (s_count.mono inter_subset_left) fun b _ => (huv i b).1 let f' : α → β := fun x => ⨅ i : s, piecewise (u' i) (fun _ => (i : β)) (fun _ => (⊤ : β)) x - have f'_meas : Measurable f' := by - apply measurable_iInf - exact fun i => Measurable.piecewise (u'_meas i) measurable_const measurable_const + have f'_meas : Measurable f' := by fun_prop (disch := aesop) let t := ⋃ (p : s) (q : ↥(s ∩ Ioi p)), u' p ∩ v p q have μt : μ t ≤ 0 := calc diff --git a/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean b/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean index b7dc71cd9a9e8..71c2449fa8236 100644 --- a/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean +++ b/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean @@ -1332,7 +1332,7 @@ theorem ae_bdd_liminf_atTop_rpow_of_eLpNorm_bdd {p : ℝ≥0∞} {f : ℕ → α have hp : p ≠ 0 := fun h => by simp [h] at hp0 have hp' : p ≠ ∞ := fun h => by simp [h] at hp0 refine - ae_lt_top (measurable_liminf fun n => (hfmeas n).nnnorm.coe_nnreal_ennreal.pow_const p.toReal) + ae_lt_top (.liminf fun n => (hfmeas n).nnnorm.coe_nnreal_ennreal.pow_const p.toReal) (lt_of_le_of_lt (lintegral_liminf_le fun n => (hfmeas n).nnnorm.coe_nnreal_ennreal.pow_const p.toReal) (lt_of_le_of_lt ?_ diff --git a/Mathlib/MeasureTheory/Integral/Lebesgue.lean b/Mathlib/MeasureTheory/Integral/Lebesgue.lean index 511c11d6a00dd..b30251af343f7 100644 --- a/Mathlib/MeasureTheory/Integral/Lebesgue.lean +++ b/Mathlib/MeasureTheory/Integral/Lebesgue.lean @@ -171,7 +171,7 @@ theorem exists_measurable_le_lintegral_eq (f : α → ℝ≥0∞) : (hLf n).2 choose g hgm hgf hLg using this refine - ⟨fun x => ⨆ n, g n x, measurable_iSup hgm, fun x => iSup_le fun n => hgf n x, le_antisymm ?_ ?_⟩ + ⟨fun x => ⨆ n, g n x, .iSup hgm, fun x => iSup_le fun n => hgf n x, le_antisymm ?_ ?_⟩ · refine le_of_tendsto' hL_tendsto fun n => (hLg n).le.trans <| lintegral_mono fun x => ?_ exact le_iSup (fun n => g n x) n · exact lintegral_mono fun x => iSup_le fun n => hgf n x @@ -1034,7 +1034,7 @@ theorem lintegral_iInf_ae {f : ℕ → α → ℝ≥0∞} (h_meas : ∀ n, Measu show ∫⁻ a, f 0 a ∂μ - ∫⁻ a, ⨅ n, f n a ∂μ = ∫⁻ a, f 0 a ∂μ - ⨅ n, ∫⁻ a, f n a ∂μ from calc ∫⁻ a, f 0 a ∂μ - ∫⁻ a, ⨅ n, f n a ∂μ = ∫⁻ a, f 0 a - ⨅ n, f n a ∂μ := - (lintegral_sub (measurable_iInf h_meas) + (lintegral_sub (.iInf h_meas) (ne_top_of_le_ne_top h_fin <| lintegral_mono fun _ => iInf_le _ _) (ae_of_all _ fun _ => iInf_le _ _)).symm _ = ∫⁻ a, ⨆ n, f 0 a - f n a ∂μ := congr rfl (funext fun _ => ENNReal.sub_iInf) @@ -1113,7 +1113,7 @@ theorem lintegral_liminf_le' {f : ℕ → α → ℝ≥0∞} (h_meas : ∀ n, AE ∫⁻ a, liminf (fun n => f n a) atTop ∂μ = ∫⁻ a, ⨆ n : ℕ, ⨅ i ≥ n, f i a ∂μ := by simp only [liminf_eq_iSup_iInf_of_nat] _ = ⨆ n : ℕ, ∫⁻ a, ⨅ i ≥ n, f i a ∂μ := - (lintegral_iSup' (fun _ => aemeasurable_biInf _ (to_countable _) (fun i _ ↦ h_meas i)) + (lintegral_iSup' (fun _ => .biInf _ (to_countable _) (fun i _ ↦ h_meas i)) (ae_of_all μ fun _ _ _ hnm => iInf_le_iInf_of_subset fun _ hi => le_trans hnm hi)) _ ≤ ⨆ n : ℕ, ⨅ i ≥ n, ∫⁻ a, f i a ∂μ := iSup_mono fun _ => le_iInf₂_lintegral _ _ = atTop.liminf fun n => ∫⁻ a, f n a ∂μ := Filter.liminf_eq_iSup_iInf_of_nat.symm @@ -1133,7 +1133,7 @@ theorem limsup_lintegral_le {f : ℕ → α → ℝ≥0∞} (g : α → ℝ≥0 _ = ∫⁻ a, ⨅ n : ℕ, ⨆ i ≥ n, f i a ∂μ := by refine (lintegral_iInf ?_ ?_ ?_).symm · intro n - exact measurable_biSup _ (to_countable _) (fun i _ ↦ hf_meas i) + exact .biSup _ (to_countable _) (fun i _ ↦ hf_meas i) · intro n m hnm a exact iSup_le_iSup_of_subset fun i hi => le_trans hnm hi · refine ne_top_of_le_ne_top h_fin (lintegral_mono_ae ?_) @@ -1837,7 +1837,7 @@ theorem exists_measurable_le_forall_setLIntegral_eq [SFinite μ] (f : α → ℝ -- Without loss of generality, `μ` is a finite measure. wlog h : IsFiniteMeasure μ generalizing μ · choose g hgm hgle hgint using fun n ↦ @this (sfiniteSeq μ n) _ inferInstance - refine ⟨fun x ↦ ⨆ n, g n x, measurable_iSup hgm, fun x ↦ iSup_le (hgle · x), fun s ↦ ?_⟩ + refine ⟨fun x ↦ ⨆ n, g n x, .iSup hgm, fun x ↦ iSup_le (hgle · x), fun s ↦ ?_⟩ rw [← sum_sfiniteSeq μ, Measure.restrict_sum_of_countable, lintegral_sum_measure, lintegral_sum_measure] exact ENNReal.tsum_le_tsum fun n ↦ (hgint n s).trans (lintegral_mono fun x ↦ le_iSup (g · x) _) diff --git a/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean b/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean index b252ea984d167..fa14f94db23e7 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean @@ -287,7 +287,7 @@ protected theorem MeasurableSet.preimage {t : Set β} (ht : MeasurableSet t) (hf MeasurableSet (f ⁻¹' t) := hf ht -@[measurability] +@[measurability, fun_prop] protected theorem Measurable.piecewise {_ : DecidablePred (· ∈ s)} (hs : MeasurableSet s) (hf : Measurable f) (hg : Measurable g) : Measurable (piecewise s f g) := by intro t ht diff --git a/Mathlib/MeasureTheory/Measure/GiryMonad.lean b/Mathlib/MeasureTheory/Measure/GiryMonad.lean index 5f92026d9e1e7..4b954b6e64486 100644 --- a/Mathlib/MeasureTheory/Measure/GiryMonad.lean +++ b/Mathlib/MeasureTheory/Measure/GiryMonad.lean @@ -78,7 +78,7 @@ theorem measurable_dirac : Measurable (Measure.dirac : α → Measure α) := by theorem measurable_lintegral {f : α → ℝ≥0∞} (hf : Measurable f) : Measurable fun μ : Measure α => ∫⁻ x, f x ∂μ := by simp only [lintegral_eq_iSup_eapprox_lintegral, hf, SimpleFunc.lintegral] - refine measurable_iSup fun n => Finset.measurable_sum _ fun i _ => ?_ + refine .iSup fun n => Finset.measurable_sum _ fun i _ => ?_ refine Measurable.const_mul ?_ _ exact measurable_coe ((SimpleFunc.eapprox f n).measurableSet_preimage _) diff --git a/Mathlib/Probability/Kernel/Disintegration/CDFToKernel.lean b/Mathlib/Probability/Kernel/Disintegration/CDFToKernel.lean index cf2267887837b..24c15a8c5b0bb 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CDFToKernel.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CDFToKernel.lean @@ -333,7 +333,7 @@ lemma IsRatCondKernelCDFAux.integrable_iInf_rat_gt (hf : IsRatCondKernelCDFAux f [IsFiniteKernel ν] (a : α) (q : ℚ) : Integrable (fun t ↦ ⨅ r : Ioi q, f (a, t) r) (ν a) := by rw [← memℒp_one_iff_integrable] - refine ⟨(measurable_iInf fun i ↦ hf.measurable_right a _).aestronglyMeasurable, ?_⟩ + refine ⟨(Measurable.iInf fun i ↦ hf.measurable_right a _).aestronglyMeasurable, ?_⟩ refine (?_ : _ ≤ (ν a univ : ℝ≥0∞)).trans_lt (measure_lt_top _ _) refine (eLpNorm_le_of_ae_bound (C := 1) ?_).trans (by simp) filter_upwards [hf.bddBelow_range a, hf.nonneg a, hf.le_one a] diff --git a/Mathlib/Probability/Kernel/Disintegration/Density.lean b/Mathlib/Probability/Kernel/Disintegration/Density.lean index 61538fa7f84ae..e66eaa2f6fe3f 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Density.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Density.lean @@ -484,7 +484,7 @@ lemma tendsto_m_density (hκν : fst κ ≤ ν) (a : α) [IsFiniteKernel ν] lemma measurable_density (κ : Kernel α (γ × β)) (ν : Kernel α γ) {s : Set β} (hs : MeasurableSet s) : Measurable (fun (p : α × γ) ↦ density κ ν p.1 p.2 s) := - measurable_limsup (fun n ↦ measurable_densityProcess κ ν n hs) + .limsup (fun n ↦ measurable_densityProcess κ ν n hs) lemma measurable_density_left (κ : Kernel α (γ × β)) (ν : Kernel α γ) (x : γ) {s : Set β} (hs : MeasurableSet s) : diff --git a/Mathlib/Probability/Kernel/Disintegration/MeasurableStieltjes.lean b/Mathlib/Probability/Kernel/Disintegration/MeasurableStieltjes.lean index 9ecc3f1b58b8d..3dfd2fa43ee37 100644 --- a/Mathlib/Probability/Kernel/Disintegration/MeasurableStieltjes.lean +++ b/Mathlib/Probability/Kernel/Disintegration/MeasurableStieltjes.lean @@ -84,7 +84,7 @@ lemma measurableSet_isRatStieltjesPoint [MeasurableSpace α] (hf : Measurable f) have h4 : MeasurableSet {a | ∀ t : ℚ, ⨅ r : Ioi t, f a r = f a t} := by rw [Set.setOf_forall] refine MeasurableSet.iInter (fun q ↦ ?_) - exact measurableSet_eq_fun (measurable_iInf fun _ ↦ hf.eval) hf.eval + exact measurableSet_eq_fun (.iInf fun _ ↦ hf.eval) hf.eval suffices {a | IsRatStieltjesPoint f a} = ({a | Monotone (f a)} ∩ {a | Tendsto (f a) atTop (𝓝 1)} ∩ {a | Tendsto (f a) atBot (𝓝 0)} ∩ {a | ∀ t : ℚ, ⨅ r : Ioi t, f a r = f a t}) by @@ -388,7 +388,7 @@ lemma IsMeasurableRatCDF.measurable_stieltjesFunction (x : ℝ) : congr with q rw [stieltjesFunction_eq] rw [this] - exact measurable_iInf (fun q ↦ hf.measurable.eval) + exact .iInf (fun q ↦ hf.measurable.eval) lemma IsMeasurableRatCDF.stronglyMeasurable_stieltjesFunction (x : ℝ) : StronglyMeasurable fun a ↦ hf.stieltjesFunction a x := diff --git a/Mathlib/Probability/Kernel/MeasurableIntegral.lean b/Mathlib/Probability/Kernel/MeasurableIntegral.lean index b35f6f0a1fcab..2d5fff10ca258 100644 --- a/Mathlib/Probability/Kernel/MeasurableIntegral.lean +++ b/Mathlib/Probability/Kernel/MeasurableIntegral.lean @@ -152,7 +152,7 @@ theorem _root_.Measurable.lintegral_kernel_prod_right {f : α → β → ℝ≥0 · exact fun n => (F n).measurable.comp measurable_prod_mk_left · exact fun i j hij b => SimpleFunc.monotone_eapprox (uncurry f) hij _ simp_rw [this] - refine measurable_iSup fun n => ?_ + refine .iSup fun n => ?_ refine SimpleFunc.induction (P := fun f => Measurable (fun (a : α) => ∫⁻ (b : β), f (a, b) ∂κ a)) ?_ ?_ (F n) · intro c t ht diff --git a/Mathlib/Probability/Martingale/Upcrossing.lean b/Mathlib/Probability/Martingale/Upcrossing.lean index 3550483556921..f4b820cc62dbc 100644 --- a/Mathlib/Probability/Martingale/Upcrossing.lean +++ b/Mathlib/Probability/Martingale/Upcrossing.lean @@ -762,7 +762,7 @@ noncomputable def upcrossings [Preorder ι] [OrderBot ι] [InfSet ι] (a b : ℝ theorem Adapted.measurable_upcrossings (hf : Adapted ℱ f) (hab : a < b) : Measurable (upcrossings a b f) := - measurable_iSup fun _ => measurable_from_top.comp (hf.measurable_upcrossingsBefore hab) + .iSup fun _ => measurable_from_top.comp (hf.measurable_upcrossingsBefore hab) theorem upcrossings_lt_top_iff : upcrossings a b f ω < ∞ ↔ ∃ k, ∀ N, upcrossingsBefore a b f N ω ≤ k := by From b2cf5147719e6c062e20790f1e7dce03948ead3a Mon Sep 17 00:00:00 2001 From: Yakov Pechersky Date: Tue, 22 Oct 2024 08:37:15 +0000 Subject: [PATCH 260/505] chore(RingTheory/PrincipalIdealDomain): factor out `IsPrincipal.map`, use RingHomClass (#18010) --- Mathlib/RingTheory/PrincipalIdealDomain.lean | 33 +++++++++++++------- 1 file changed, 21 insertions(+), 12 deletions(-) diff --git a/Mathlib/RingTheory/PrincipalIdealDomain.lean b/Mathlib/RingTheory/PrincipalIdealDomain.lean index 40c180a84813b..6e99a6bcbaf3f 100644 --- a/Mathlib/RingTheory/PrincipalIdealDomain.lean +++ b/Mathlib/RingTheory/PrincipalIdealDomain.lean @@ -348,23 +348,32 @@ section Surjective open Submodule -variable {S N : Type*} [Ring R] [AddCommGroup M] [AddCommGroup N] [Ring S] -variable [Module R M] [Module R N] +variable {S N F : Type*} [Ring R] [AddCommGroup M] [AddCommGroup N] [Ring S] +variable [Module R M] [Module R N] [FunLike F R S] [RingHomClass F R S] + +theorem Submodule.IsPrincipal.map (f : M →ₗ[R] N) {S : Submodule R M} + (hI : IsPrincipal S) : IsPrincipal (map f S) := + ⟨⟨f (IsPrincipal.generator S), by + rw [← Set.image_singleton, ← map_span, span_singleton_generator]⟩⟩ theorem Submodule.IsPrincipal.of_comap (f : M →ₗ[R] N) (hf : Function.Surjective f) - (S : Submodule R N) [hI : IsPrincipal (S.comap f)] : IsPrincipal S := - ⟨⟨f (IsPrincipal.generator (S.comap f)), by - rw [← Set.image_singleton, ← Submodule.map_span, IsPrincipal.span_singleton_generator, - Submodule.map_comap_eq_of_surjective hf]⟩⟩ - -theorem Ideal.IsPrincipal.of_comap (f : R →+* S) (hf : Function.Surjective f) (I : Ideal S) - [hI : IsPrincipal (I.comap f)] : IsPrincipal I := - ⟨⟨f (IsPrincipal.generator (I.comap f)), by + (S : Submodule R N) [hI : IsPrincipal (S.comap f)] : IsPrincipal S := by + rw [← Submodule.map_comap_eq_of_surjective hf S] + exact hI.map f + +theorem Submodule.IsPrincipal.map_ringHom (f : F) {I : Ideal R} + (hI : IsPrincipal I) : IsPrincipal (Ideal.map f I) := + ⟨⟨f (IsPrincipal.generator I), by rw [Ideal.submodule_span_eq, ← Set.image_singleton, ← Ideal.map_span, - Ideal.span_singleton_generator, Ideal.map_comap_of_surjective f hf]⟩⟩ + Ideal.span_singleton_generator]⟩⟩ + +theorem Ideal.IsPrincipal.of_comap (f : F) (hf : Function.Surjective f) (I : Ideal S) + [hI : IsPrincipal (I.comap f)] : IsPrincipal I := by + rw [← map_comap_of_surjective f hf I] + exact hI.map_ringHom f /-- The surjective image of a principal ideal ring is again a principal ideal ring. -/ -theorem IsPrincipalIdealRing.of_surjective [IsPrincipalIdealRing R] (f : R →+* S) +theorem IsPrincipalIdealRing.of_surjective [IsPrincipalIdealRing R] (f : F) (hf : Function.Surjective f) : IsPrincipalIdealRing S := ⟨fun I => Ideal.IsPrincipal.of_comap f hf I⟩ From 5eb330bc038e8d4d8af4f844fbe9a87b896e89aa Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Tue, 22 Oct 2024 08:37:17 +0000 Subject: [PATCH 261/505] chore(Algebra/PUnitInstances): split off `GCDMonoid PUnit` instance (#18014) This instance does not seem to be used anywhere in Mathlib and splitting it off should allow us to shave quite a few imports from low down in the hierarchy. --- Mathlib.lean | 1 + Mathlib/Algebra/Group/AddChar.lean | 1 + Mathlib/Algebra/PUnitInstances/Algebra.lean | 31 +----------- Mathlib/Algebra/PUnitInstances/GCDMonoid.lean | 47 +++++++++++++++++++ Mathlib/Algebra/Star/CHSH.lean | 1 + Mathlib/RingTheory/Ideal/Basic.lean | 6 ++- 6 files changed, 55 insertions(+), 32 deletions(-) create mode 100644 Mathlib/Algebra/PUnitInstances/GCDMonoid.lean diff --git a/Mathlib.lean b/Mathlib.lean index 01695686dd757..a904678156efa 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -711,6 +711,7 @@ import Mathlib.Algebra.Order.UpperLower import Mathlib.Algebra.Order.ZeroLEOne import Mathlib.Algebra.PEmptyInstances import Mathlib.Algebra.PUnitInstances.Algebra +import Mathlib.Algebra.PUnitInstances.GCDMonoid import Mathlib.Algebra.PUnitInstances.Module import Mathlib.Algebra.PUnitInstances.Order import Mathlib.Algebra.Periodic diff --git a/Mathlib/Algebra/Group/AddChar.lean b/Mathlib/Algebra/Group/AddChar.lean index 23df07dbc7f89..23505cff4f752 100644 --- a/Mathlib/Algebra/Group/AddChar.lean +++ b/Mathlib/Algebra/Group/AddChar.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Michael Stoll. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Michael Stoll -/ +import Mathlib.Algebra.Ring.Regular import Mathlib.Logic.Equiv.TransferInstance /-! diff --git a/Mathlib/Algebra/PUnitInstances/Algebra.lean b/Mathlib/Algebra/PUnitInstances/Algebra.lean index a9860604f2c27..8a9c98a4aebdc 100644 --- a/Mathlib/Algebra/PUnitInstances/Algebra.lean +++ b/Mathlib/Algebra/PUnitInstances/Algebra.lean @@ -3,7 +3,7 @@ Copyright (c) 2019 Kenny Lau. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau -/ -import Mathlib.Algebra.GCDMonoid.Basic +import Mathlib.Algebra.Ring.Basic /-! # Instances on PUnit @@ -64,33 +64,4 @@ instance commRing : CommRing PUnit where instance cancelCommMonoidWithZero : CancelCommMonoidWithZero PUnit where --- This is too high-powered and should be split off also -instance normalizedGCDMonoid : NormalizedGCDMonoid PUnit where - gcd _ _ := unit - lcm _ _ := unit - normUnit _ := 1 - normUnit_zero := rfl - normUnit_mul := by intros; rfl - normUnit_coe_units := by intros; rfl - gcd_dvd_left _ _ := ⟨unit, by subsingleton⟩ - gcd_dvd_right _ _ := ⟨unit, by subsingleton⟩ - dvd_gcd {_ _} _ _ _ := ⟨unit, by subsingleton⟩ - gcd_mul_lcm _ _ := ⟨1, by subsingleton⟩ - lcm_zero_left := by intros; rfl - lcm_zero_right := by intros; rfl - normalize_gcd := by intros; rfl - normalize_lcm := by intros; rfl - --- Porting note (#10618): simpNF lint: simp can prove this @[simp] -theorem gcd_eq {x y : PUnit} : gcd x y = unit := - rfl - --- Porting note (#10618): simpNF lint: simp can prove this @[simp] -theorem lcm_eq {x y : PUnit} : lcm x y = unit := - rfl - -@[simp] -theorem norm_unit_eq {x : PUnit} : normUnit x = 1 := - rfl - end PUnit diff --git a/Mathlib/Algebra/PUnitInstances/GCDMonoid.lean b/Mathlib/Algebra/PUnitInstances/GCDMonoid.lean new file mode 100644 index 0000000000000..d80a17643fd74 --- /dev/null +++ b/Mathlib/Algebra/PUnitInstances/GCDMonoid.lean @@ -0,0 +1,47 @@ +/- +Copyright (c) 2019 Kenny Lau. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kenny Lau +-/ +import Mathlib.Algebra.GCDMonoid.Basic +import Mathlib.Algebra.PUnitInstances.Algebra + +/-! +# Instances on PUnit + +This file collects facts about algebraic structures on the one-element type, e.g. that it is a +commutative ring. +-/ + +namespace PUnit + +-- This is too high-powered and should be split off also +instance normalizedGCDMonoid : NormalizedGCDMonoid PUnit where + gcd _ _ := unit + lcm _ _ := unit + normUnit _ := 1 + normUnit_zero := rfl + normUnit_mul := by intros; rfl + normUnit_coe_units := by intros; rfl + gcd_dvd_left _ _ := ⟨unit, by subsingleton⟩ + gcd_dvd_right _ _ := ⟨unit, by subsingleton⟩ + dvd_gcd {_ _} _ _ _ := ⟨unit, by subsingleton⟩ + gcd_mul_lcm _ _ := ⟨1, by subsingleton⟩ + lcm_zero_left := by intros; rfl + lcm_zero_right := by intros; rfl + normalize_gcd := by intros; rfl + normalize_lcm := by intros; rfl + +-- Porting note (#10618): simpNF lint: simp can prove this @[simp] +theorem gcd_eq {x y : PUnit} : gcd x y = unit := + rfl + +-- Porting note (#10618): simpNF lint: simp can prove this @[simp] +theorem lcm_eq {x y : PUnit} : lcm x y = unit := + rfl + +@[simp] +theorem norm_unit_eq {x : PUnit} : normUnit x = 1 := + rfl + +end PUnit diff --git a/Mathlib/Algebra/Star/CHSH.lean b/Mathlib/Algebra/Star/CHSH.lean index 26fc0b59288bb..eedcf9c9f9801 100644 --- a/Mathlib/Algebra/Star/CHSH.lean +++ b/Mathlib/Algebra/Star/CHSH.lean @@ -5,6 +5,7 @@ Authors: Kim Morrison -/ import Mathlib.Algebra.CharP.Invertible import Mathlib.Algebra.Order.Star.Basic +import Mathlib.Algebra.Ring.Regular import Mathlib.Data.Real.Sqrt import Mathlib.Tactic.Polyrith diff --git a/Mathlib/RingTheory/Ideal/Basic.lean b/Mathlib/RingTheory/Ideal/Basic.lean index 1db54f75b9e6e..e2c163dc62df6 100644 --- a/Mathlib/RingTheory/Ideal/Basic.lean +++ b/Mathlib/RingTheory/Ideal/Basic.lean @@ -3,11 +3,13 @@ Copyright (c) 2018 Kenny Lau. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau, Chris Hughes, Mario Carneiro -/ -import Mathlib.Tactic.FinCases +import Mathlib.Algebra.Associated.Basic +import Mathlib.Algebra.Field.IsField +import Mathlib.Algebra.Ring.Regular import Mathlib.Data.Nat.Choose.Sum import Mathlib.LinearAlgebra.Finsupp -import Mathlib.Algebra.Field.IsField import Mathlib.Tactic.Abel +import Mathlib.Tactic.FinCases /-! From 1460f4a06f36438984bb9f5b608d5fcd75fb8db6 Mon Sep 17 00:00:00 2001 From: Scott Carnahan <128885296+ScottCarnahan@users.noreply.github.com> Date: Tue, 22 Oct 2024 09:16:16 +0000 Subject: [PATCH 262/505] feat(RingTheory/LaurentSeries): composition of Hasse Derivatives (#16106) This PR proves basic properties of Hasse derivatives of Laurent series including composition, and introduces the derivative as a linear map. --- Mathlib/RingTheory/HahnSeries/Addition.lean | 6 ++ Mathlib/RingTheory/LaurentSeries.lean | 63 +++++++++++++++++++-- 2 files changed, 65 insertions(+), 4 deletions(-) diff --git a/Mathlib/RingTheory/HahnSeries/Addition.lean b/Mathlib/RingTheory/HahnSeries/Addition.lean index 255b47ebace23..a7e06d6663a72 100644 --- a/Mathlib/RingTheory/HahnSeries/Addition.lean +++ b/Mathlib/RingTheory/HahnSeries/Addition.lean @@ -66,6 +66,12 @@ theorem add_coeff' {x y : HahnSeries Γ R} : (x + y).coeff = x.coeff + y.coeff : theorem add_coeff {x y : HahnSeries Γ R} {a : Γ} : (x + y).coeff a = x.coeff a + y.coeff a := rfl +@[simp] +theorem nsmul_coeff {x : HahnSeries Γ R} {n : ℕ} : (n • x).coeff = n • x.coeff := by + induction n with + | zero => simp + | succ n ih => simp [add_nsmul, ih] + @[simp] protected lemma map_add [AddMonoid S] (f : R →+ S) {x y : HahnSeries Γ R} : ((x + y).map f : HahnSeries Γ S) = x.map f + y.map f := by diff --git a/Mathlib/RingTheory/LaurentSeries.lean b/Mathlib/RingTheory/LaurentSeries.lean index e9cc56ec5ec51..c588752360d21 100644 --- a/Mathlib/RingTheory/LaurentSeries.lean +++ b/Mathlib/RingTheory/LaurentSeries.lean @@ -4,12 +4,11 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Aaron Anderson, María Inés de Frutos-Fernández, Filippo A. E. Nuccio -/ import Mathlib.Data.Int.Interval +import Mathlib.FieldTheory.RatFunc.AsPolynomial import Mathlib.RingTheory.Binomial -import Mathlib.RingTheory.DedekindDomain.Basic import Mathlib.RingTheory.HahnSeries.PowerSeries import Mathlib.RingTheory.HahnSeries.Summable import Mathlib.RingTheory.PowerSeries.Inverse -import Mathlib.FieldTheory.RatFunc.AsPolynomial import Mathlib.RingTheory.Localization.FractionRing import Mathlib.Topology.UniformSpace.Cauchy @@ -86,7 +85,6 @@ end section HasseDeriv /-- The Hasse derivative of Laurent series, as a linear map. -/ -@[simps] def hasseDeriv (R : Type*) {V : Type*} [AddCommGroup V] [Semiring R] [Module R V] (k : ℕ) : V⸨X⸩ →ₗ[R] V⸨X⸩ where toFun f := HahnSeries.ofSuppBddBelow (fun (n : ℤ) => (Ring.choose (n + k) k) • f.coeff (n + k)) @@ -101,10 +99,67 @@ def hasseDeriv (R : Type*) {V : Type*} [AddCommGroup V] [Semiring R] [Module R V variable [Semiring R] {V : Type*} [AddCommGroup V] [Module R V] -theorem hasseDeriv_coeff (k : ℕ) (f : V⸨X⸩) (n : ℤ) : +@[simp] +theorem hasseDeriv_coeff (k : ℕ) (f : LaurentSeries V) (n : ℤ) : (hasseDeriv R k f).coeff n = Ring.choose (n + k) k • f.coeff (n + k) := rfl +@[simp] +theorem hasseDeriv_zero : hasseDeriv R 0 = LinearMap.id (M := LaurentSeries V) := by + ext f n + simp + +theorem hasseDeriv_single_add (k : ℕ) (n : ℤ) (x : V) : + hasseDeriv R k (single (n + k) x) = single n ((Ring.choose (n + k) k) • x) := by + ext m + dsimp only [hasseDeriv_coeff] + by_cases h : m = n + · simp [h] + · simp [h, show m + k ≠ n + k by omega] + +@[simp] +theorem hasseDeriv_single (k : ℕ) (n : ℤ) (x : V) : + hasseDeriv R k (single n x) = single (n - k) ((Ring.choose n k) • x) := by + rw [← Int.sub_add_cancel n k, hasseDeriv_single_add, Int.sub_add_cancel n k] + +theorem hasseDeriv_comp_coeff (k l : ℕ) (f : LaurentSeries V) (n : ℤ) : + (hasseDeriv R k (hasseDeriv R l f)).coeff n = + ((Nat.choose (k + l) k) • hasseDeriv R (k + l) f).coeff n := by + rw [nsmul_coeff] + simp only [hasseDeriv_coeff, Pi.smul_apply, Nat.cast_add] + rw [smul_smul, mul_comm, ← Ring.choose_add_smul_choose (n + k), add_assoc, Nat.choose_symm_add, + smul_assoc] + +@[simp] +theorem hasseDeriv_comp (k l : ℕ) (f : LaurentSeries V) : + hasseDeriv R k (hasseDeriv R l f) = (k + l).choose k • hasseDeriv R (k + l) f := by + ext n + simp [hasseDeriv_comp_coeff k l f n] + +/-- The derivative of a Laurent series. -/ +def derivative (R : Type*) {V : Type*} [AddCommGroup V] [Semiring R] [Module R V] : + LaurentSeries V →ₗ[R] LaurentSeries V := + hasseDeriv R 1 + +@[simp] +theorem derivative_apply (f : LaurentSeries V) : derivative R f = hasseDeriv R 1 f := by + exact rfl + +theorem derivative_iterate (k : ℕ) (f : LaurentSeries V) : + (derivative R)^[k] f = k.factorial • (hasseDeriv R k f) := by + ext n + induction k generalizing f with + | zero => simp + | succ k ih => + rw [Function.iterate_succ, Function.comp_apply, ih, derivative_apply, hasseDeriv_comp, + Nat.choose_symm_add, Nat.choose_one_right, Nat.factorial, mul_nsmul] + +@[simp] +theorem derivative_iterate_coeff (k : ℕ) (f : LaurentSeries V) (n : ℤ) : + ((derivative R)^[k] f).coeff n = (descPochhammer ℤ k).smeval (n + k) • f.coeff (n + k) := by + rw [derivative_iterate, nsmul_coeff, Pi.smul_apply, hasseDeriv_coeff, + Ring.descPochhammer_eq_factorial_smul_choose, smul_assoc] + end HasseDeriv section Semiring From a7d3e7420fd596e31e21d2f1a03b6e48a7d6c9da Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Tue, 22 Oct 2024 09:16:17 +0000 Subject: [PATCH 263/505] feat: Algebra.IsAlgebraic.tower_bot (#17893) Co-authored-by: FR --- Mathlib/RingTheory/Algebraic.lean | 12 ++++++++++++ 1 file changed, 12 insertions(+) diff --git a/Mathlib/RingTheory/Algebraic.lean b/Mathlib/RingTheory/Algebraic.lean index 9942e6c347371..89dc73d0fbe22 100644 --- a/Mathlib/RingTheory/Algebraic.lean +++ b/Mathlib/RingTheory/Algebraic.lean @@ -324,6 +324,12 @@ theorem Algebra.IsAlgebraic.tower_top_of_injective (hinj : Function.Injective (a [Algebra.IsAlgebraic R A] : Algebra.IsAlgebraic S A := ⟨fun _ ↦ _root_.IsAlgebraic.tower_top_of_injective hinj (Algebra.IsAlgebraic.isAlgebraic _)⟩ +theorem Algebra.IsAlgebraic.tower_bot_of_injective [Algebra.IsAlgebraic R A] + (hinj : Function.Injective (algebraMap S A)) : + Algebra.IsAlgebraic R S where + isAlgebraic x := by + simpa [isAlgebraic_algebraMap_iff hinj] using isAlgebraic (R := R) (A := A) (algebraMap _ _ x) + end CommRing section Field @@ -352,6 +358,12 @@ variable (A) instance Algebra.IsAlgebraic.of_finite [FiniteDimensional K A] : Algebra.IsAlgebraic K A := (IsIntegral.of_finite K A).isAlgebraic +theorem Algebra.IsAlgebraic.tower_bot (K L A : Type*) [CommRing K] [Field L] [Ring A] + [Algebra K L] [Algebra L A] [Algebra K A] [IsScalarTower K L A] + [Nontrivial A] [Algebra.IsAlgebraic K A] : + Algebra.IsAlgebraic K L := + tower_bot_of_injective (algebraMap L A).injective + end Field end Ring From b5879a46b78fb5ba12ecf5c64042a4dd7f2a4cf2 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Tue, 22 Oct 2024 09:16:19 +0000 Subject: [PATCH 264/505] feat: Multiset.pow_smul_esymm (#17896) Co-authored-by: FR --- Mathlib/RingTheory/MvPolynomial/Symmetric/Defs.lean | 9 +++++++++ 1 file changed, 9 insertions(+) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric/Defs.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric/Defs.lean index 691265878f204..2b4ebd09a9114 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric/Defs.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric/Defs.lean @@ -71,6 +71,15 @@ theorem _root_.Finset.esymm_map_val {σ} (f : σ → R) (s : Finset σ) (n : ℕ simp only [esymm, powersetCard_map, ← Finset.map_val_val_powersetCard, map_map] rfl +lemma pow_smul_esymm {S : Type*} [Monoid S] [DistribMulAction S R] [IsScalarTower S R R] + [SMulCommClass S R R] (s : S) (n : ℕ) (m : Multiset R) : + s ^ n • m.esymm n = (m.map (s • ·)).esymm n := by + rw [esymm, smul_sum, map_map] + trans ((powersetCard n m).map (fun x : Multiset R ↦ s ^ card x • x.prod)).sum + · refine congr_arg _ (map_congr rfl (fun x hx ↦ ?_)) + rw [Function.comp_apply, (mem_powersetCard.1 hx).2] + · simp_rw [smul_prod, esymm, powersetCard_map, map_map, Function.comp_def] + end Multiset namespace MvPolynomial From d65220b7f354f34ba74f8919677fc66df3e65009 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 22 Oct 2024 09:16:20 +0000 Subject: [PATCH 265/505] feat: `Coalgebra (DFinsupp _)` (#17972) This is pretty much a copy-paste of the `Finsupp` version, and can be used to get the instance on `DirectSum` for free. --- Mathlib/LinearAlgebra/DFinsupp.lean | 8 +++ Mathlib/RingTheory/Coalgebra/Basic.lean | 66 ++++++++++++++++++++++++- 2 files changed, 72 insertions(+), 2 deletions(-) diff --git a/Mathlib/LinearAlgebra/DFinsupp.lean b/Mathlib/LinearAlgebra/DFinsupp.lean index 292f5f03db793..52dc2d02af921 100644 --- a/Mathlib/LinearAlgebra/DFinsupp.lean +++ b/Mathlib/LinearAlgebra/DFinsupp.lean @@ -89,6 +89,14 @@ def lapply (i : ι) : (Π₀ i, M i) →ₗ[R] M i where theorem lapply_apply (i : ι) (f : Π₀ i, M i) : (lapply i : (Π₀ i, M i) →ₗ[R] _) f = f i := rfl +@[simp] +theorem lapply_comp_lsingle_same [DecidableEq ι] (i : ι) : + lapply i ∘ₗ lsingle i = (.id : M i →ₗ[R] M i) := by ext; simp + +@[simp] +theorem lapply_comp_lsingle_of_ne [DecidableEq ι] (i i' : ι) (h : i ≠ i') : + lapply i ∘ₗ lsingle i' = (0 : M i' →ₗ[R] M i) := by ext; simp [h.symm] + section Lsum -- Porting note: Unclear how true these docstrings are in lean 4 diff --git a/Mathlib/RingTheory/Coalgebra/Basic.lean b/Mathlib/RingTheory/Coalgebra/Basic.lean index 643086a03c0e1..bb611315a00d3 100644 --- a/Mathlib/RingTheory/Coalgebra/Basic.lean +++ b/Mathlib/RingTheory/Coalgebra/Basic.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Ali Ramsey, Eric Wieser -/ import Mathlib.Algebra.Algebra.Bilinear -import Mathlib.LinearAlgebra.Finsupp +import Mathlib.LinearAlgebra.DFinsupp import Mathlib.LinearAlgebra.Prod import Mathlib.LinearAlgebra.TensorProduct.Finiteness @@ -15,7 +15,7 @@ In this file we define `Coalgebra`, and provide instances for: * Commutative semirings: `CommSemiring.toCoalgebra` * Binary products: `Prod.instCoalgebra` -* Finitely supported functions: `Finsupp.instCoalgebra` +* Finitely supported functions: `DFinsupp.instCoalgebra`, `Finsupp.instCoalgebra` ## References @@ -259,6 +259,68 @@ instance instCoalgebra : Coalgebra R (A × B) where end Prod +namespace DFinsupp +variable (R : Type u) (ι : Type v) (A : ι → Type w) +variable [DecidableEq ι] +variable [CommSemiring R] [∀ i, AddCommMonoid (A i)] [∀ i, Module R (A i)] [∀ i, Coalgebra R (A i)] + +open LinearMap + +instance instCoalgebraStruct : CoalgebraStruct R (Π₀ i, A i) where + comul := DFinsupp.lsum R fun i => + TensorProduct.map (DFinsupp.lsingle i) (DFinsupp.lsingle i) ∘ₗ comul + counit := DFinsupp.lsum R fun _ => counit + +@[simp] +theorem comul_single (i : ι) (a : A i) : + comul (R := R) (DFinsupp.single i a) = + (TensorProduct.map (DFinsupp.lsingle i) (DFinsupp.lsingle i) : _ →ₗ[R] _) (comul a) := + lsum_single _ _ _ _ + +@[simp] +theorem counit_single (i : ι) (a : A i) : counit (DFinsupp.single i a) = counit (R := R) a := + lsum_single _ _ _ _ + +theorem comul_comp_lsingle (i : ι) : + comul ∘ₗ (lsingle i : A i →ₗ[R] _) = TensorProduct.map (lsingle i) (lsingle i) ∘ₗ comul := by + ext; simp + +theorem comul_comp_lapply (i : ι) : + comul ∘ₗ (lapply i : _ →ₗ[R] A i) = TensorProduct.map (lapply i) (lapply i) ∘ₗ comul := by + ext j : 1 + conv_rhs => rw [comp_assoc, comul_comp_lsingle, ← comp_assoc, ← TensorProduct.map_comp] + obtain rfl | hij := eq_or_ne i j + · rw [comp_assoc, lapply_comp_lsingle_same, comp_id, TensorProduct.map_id, id_comp] + · rw [comp_assoc, lapply_comp_lsingle_of_ne _ _ hij, comp_zero, TensorProduct.map_zero_left, + zero_comp] + +@[simp] theorem counit_comp_lsingle (i : ι) : counit ∘ₗ (lsingle i : A i →ₗ[R] _) = counit := by + ext; simp + +/-- The `R`-module whose elements are dependent functions `(i : ι) → A i` which are zero on all but +finitely many elements of `ι` has a coalgebra structure. + +The coproduct `Δ` is given by `Δ(fᵢ a) = fᵢ a₁ ⊗ fᵢ a₂` where `Δ(a) = a₁ ⊗ a₂` and the counit `ε` +by `ε(fᵢ a) = ε(a)`, where `fᵢ a` is the function sending `i` to `a` and all other elements of `ι` +to zero. -/ +instance instCoalgebra : Coalgebra R (Π₀ i, A i) where + rTensor_counit_comp_comul := by + ext : 1 + rw [comp_assoc, comul_comp_lsingle, ← comp_assoc, rTensor_comp_map, counit_comp_lsingle, + ← lTensor_comp_rTensor, comp_assoc, rTensor_counit_comp_comul, lTensor_comp_mk] + lTensor_counit_comp_comul := by + ext : 1 + rw [comp_assoc, comul_comp_lsingle, ← comp_assoc, lTensor_comp_map, counit_comp_lsingle, + ← rTensor_comp_lTensor, comp_assoc, lTensor_counit_comp_comul, rTensor_comp_flip_mk] + coassoc := by + ext i : 1 + simp_rw [comp_assoc, comul_comp_lsingle, ← comp_assoc, lTensor_comp_map, comul_comp_lsingle, + comp_assoc, ← comp_assoc comul, rTensor_comp_map, comul_comp_lsingle, ← map_comp_rTensor, + ← map_comp_lTensor, comp_assoc, ← coassoc, ← comp_assoc comul, ← comp_assoc, + TensorProduct.map_map_comp_assoc_eq] + +end DFinsupp + namespace Finsupp variable (R : Type u) (ι : Type v) (A : Type w) variable [CommSemiring R] [AddCommMonoid A] [Module R A] [Coalgebra R A] From d713728520fc030b2248b760f99251fd0f251a5b Mon Sep 17 00:00:00 2001 From: Fabrizio Barroero Date: Tue, 22 Oct 2024 10:36:30 +0000 Subject: [PATCH 266/505] doc(NumberTheory/Ostrowski): tweak documentation (#18028) This PR modifies the documentation of the file in the hope of improving it. --- Mathlib/NumberTheory/Ostrowski.lean | 19 +++++++++++-------- 1 file changed, 11 insertions(+), 8 deletions(-) diff --git a/Mathlib/NumberTheory/Ostrowski.lean b/Mathlib/NumberTheory/Ostrowski.lean index 98af768e8637a..892af836b3bec 100644 --- a/Mathlib/NumberTheory/Ostrowski.lean +++ b/Mathlib/NumberTheory/Ostrowski.lean @@ -244,7 +244,7 @@ lemma exists_pos_mulRingNorm_eq_pow_neg : ∃ t : ℝ, 0 < t ∧ f p = p ^ (-t) refine (rpow_logb (mod_cast pprime.pos) ?_ hp0).symm simp only [ne_eq, Nat.cast_eq_one,Nat.Prime.ne_one pprime, not_false_eq_true] -/-! ## Non-archimedean case: end goal -/ +-- ## Non-archimedean case: end goal include hf_nontriv bdd in /-- If `f` is bounded and not trivial, then it is equivalent to a p-adic absolute value. -/ @@ -289,7 +289,10 @@ end Non_archimedean section Archimedean -/-! ## Archimedean case -/ +/-! ## Archimedean case + +Every unbounded absolute value is equivalent to the standard absolute value +-/ /-- The usual absolute value on ℚ. -/ def mulRingNorm_real : MulRingNorm ℚ := @@ -308,7 +311,7 @@ def mulRingNorm_real : MulRingNorm ℚ := simp only [Rat.cast_abs] rfl -/-! ## Preliminary result -/ +-- ## Preliminary result /-- Given an two integers `n, m` with `m > 1` the mulRingNorm of `n` is bounded by `m + m * f m + m * (f m) ^ 2 + ... + m * (f m) ^ d` where `d` is the number of digits of the @@ -336,7 +339,7 @@ lemma mulRingNorm_apply_le_sum_digits (n : ℕ) {m : ℕ} (hm : 1 < m) : simp only [zero_le, zero_add, tsub_zero, true_and] at hia exact (hcoef (List.mem_iff_get.mpr ⟨⟨i, hia.1⟩, hia.2.symm⟩)).le -/-! ## Step 1: if f is a MulRingNorm and f n > 1 for some natural n, then f n > 1 for all n ≥ 2 -/ +-- ## Step 1: if f is a MulRingNorm and f n > 1 for some natural n, then f n > 1 for all n ≥ 2 /-- If `f n > 1` for some `n` then `f n > 1` for all `n ≥ 2` -/ lemma one_lt_of_not_bounded (notbdd : ¬ ∀ n : ℕ, f n ≤ 1) {n₀ : ℕ} (hn₀ : 1 < n₀) : 1 < f n₀ := by @@ -391,7 +394,7 @@ lemma one_lt_of_not_bounded (notbdd : ¬ ∀ n : ℕ, f n ≤ 1) {n₀ : ℕ} (h tendsto_root_atTop_nhds_one (by positivity) exact hnlim.mul tendsto_nat_rpow_div -/-! ## step 2: given m,n ≥ 2 and |m|=m^s, |n|=n^t for s,t > 0, we have t ≤ s -/ +-- ## Step 2: given m,n ≥ 2 and |m|=m^s, |n|=n^t for s,t > 0, we have t ≤ s variable {m n : ℕ} (hm : 1 < m) (hn : 1 < n) (notbdd : ¬ ∀ (n : ℕ), f n ≤ 1) @@ -448,7 +451,7 @@ lemma mulRingNorm_le_mulRingNorm_pow_log : f n ≤ f m ^ logb m n := by fun b hb ↦ param_upperbound hm hn notbdd (not_eq_zero_of_lt hb)⟩)) include hm hn notbdd in -/-- Given m,n ≥ 2 and `f m = m ^ s`, `f n = n ^ t` for `s, t > 0`, we have `t ≤ s`. -/ +/-- Given `m,n ≥ 2` and `f m = m ^ s`, `f n = n ^ t` for `s, t > 0`, we have `t ≤ s`. -/ lemma le_of_mulRingNorm_eq {s t : ℝ} (hfm : f m = m ^ s) (hfn : f n = n ^ t) : t ≤ s := by have hmn : f n ≤ f m ^ Real.logb m n := mulRingNorm_le_mulRingNorm_pow_log hm hn notbdd rw [← Real.rpow_le_rpow_left_iff (x:=n) (mod_cast hn), ← hfn] @@ -462,7 +465,7 @@ private lemma symmetric_roles {s t : ℝ} (hfm : f m = m ^ s) (hfn : f n = n ^ t le_antisymm (le_of_mulRingNorm_eq hn hm notbdd hfn hfm) (le_of_mulRingNorm_eq hm hn notbdd hfm hfn) -/-! ## Archimedean case: end goal -/ +-- ## Archimedean case: end goal include notbdd in /-- If `f` is not bounded and not trivial, then it is equivalent to the standard absolute value on @@ -507,7 +510,7 @@ theorem mulRingNorm_equiv_standard_of_unbounded : MulRingNorm.equiv f mulRingNor end Archimedean -/-- Ostrowski's Theorem -/ +/-- **Ostrowski's Theorem** -/ theorem mulRingNorm_equiv_standard_or_padic (f : MulRingNorm ℚ) (hf_nontriv : f ≠ 1) : (MulRingNorm.equiv f mulRingNorm_real) ∨ ∃! p, ∃ (hp : Fact (p.Prime)), MulRingNorm.equiv f (mulRingNorm_padic p) := by From fb1cb210d29981e474127caeac76d2d28d3c598d Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 22 Oct 2024 10:52:22 +0000 Subject: [PATCH 267/505] chore(Analysis/SpecificLimits/Normed): move lemmas earlier (#18046) These lemmas can be proved earlier at no cost. This is one step towards reducing imports in this area of mathlib. --- Mathlib/Analysis/Asymptotics/Asymptotics.lean | 8 +++ .../Analysis/Normed/Field/InfiniteSum.lean | 7 +++ Mathlib/Analysis/Normed/Field/Lemmas.lean | 29 +++++++++++ Mathlib/Analysis/Normed/Group/Basic.lean | 10 +++- Mathlib/Analysis/SpecificLimits/Normed.lean | 52 ------------------- 5 files changed, 53 insertions(+), 53 deletions(-) diff --git a/Mathlib/Analysis/Asymptotics/Asymptotics.lean b/Mathlib/Analysis/Asymptotics/Asymptotics.lean index 47377b0bb7c8c..2ac86221aead4 100644 --- a/Mathlib/Analysis/Asymptotics/Asymptotics.lean +++ b/Mathlib/Analysis/Asymptotics/Asymptotics.lean @@ -2048,4 +2048,12 @@ end IsBigORev end ContinuousOn +/-- The (scalar) product of a sequence that tends to zero with a bounded one also tends to zero. -/ +lemma NormedField.tendsto_zero_smul_of_tendsto_zero_of_bounded {ι 𝕜 𝔸 : Type*} + [NormedDivisionRing 𝕜] [NormedAddCommGroup 𝔸] [Module 𝕜 𝔸] [BoundedSMul 𝕜 𝔸] {l : Filter ι} + {ε : ι → 𝕜} {f : ι → 𝔸} (hε : Tendsto ε l (𝓝 0)) (hf : IsBoundedUnder (· ≤ ·) l (norm ∘ f)) : + Tendsto (ε • f) l (𝓝 0) := by + rw [← isLittleO_one_iff 𝕜] at hε ⊢ + simpa using IsLittleO.smul_isBigO hε (hf.isBigO_const (one_ne_zero : (1 : 𝕜) ≠ 0)) + set_option linter.style.longFile 2200 diff --git a/Mathlib/Analysis/Normed/Field/InfiniteSum.lean b/Mathlib/Analysis/Normed/Field/InfiniteSum.lean index 720c5451551ec..195aa8c9adc27 100644 --- a/Mathlib/Analysis/Normed/Field/InfiniteSum.lean +++ b/Mathlib/Analysis/Normed/Field/InfiniteSum.lean @@ -164,3 +164,10 @@ theorem hasSum_sum_range_mul_of_summable_norm' {f g : ℕ → R} exact tsum_mul_tsum_eq_tsum_sum_range_of_summable_norm' hf h'f hg h'g end Nat + +lemma summable_of_absolute_convergence_real {f : ℕ → ℝ} : + (∃ r, Tendsto (fun n ↦ ∑ i ∈ range n, |f i|) atTop (𝓝 r)) → Summable f + | ⟨r, hr⟩ => by + refine .of_norm ⟨r, (hasSum_iff_tendsto_nat_of_nonneg ?_ _).2 ?_⟩ + · exact fun i ↦ norm_nonneg _ + · simpa only using hr diff --git a/Mathlib/Analysis/Normed/Field/Lemmas.lean b/Mathlib/Analysis/Normed/Field/Lemmas.lean index a2236fabdefec..1d663e0295913 100644 --- a/Mathlib/Analysis/Normed/Field/Lemmas.lean +++ b/Mathlib/Analysis/Normed/Field/Lemmas.lean @@ -320,6 +320,19 @@ example [Monoid β] (φ : β →* α) {x : β} {k : ℕ+} (h : x ^ (k : ℕ) = 1 @[simp] lemma AddChar.norm_apply {G : Type*} [AddLeftCancelMonoid G] [Finite G] (ψ : AddChar G α) (x : G) : ‖ψ x‖ = 1 := (ψ.toMonoidHom.isOfFinOrder <| isOfFinOrder_of_finite _).norm_eq_one +lemma NormedField.tendsto_norm_inverse_nhdsWithin_0_atTop : + Tendsto (fun x : α ↦ ‖x⁻¹‖) (𝓝[≠] 0) atTop := + (tendsto_inv_zero_atTop.comp tendsto_norm_zero').congr fun x ↦ (norm_inv x).symm + +lemma NormedField.tendsto_norm_zpow_nhdsWithin_0_atTop {m : ℤ} (hm : m < 0) : + Tendsto (fun x : α ↦ ‖x ^ m‖) (𝓝[≠] 0) atTop := by + obtain ⟨m, rfl⟩ := neg_surjective m + rw [neg_lt_zero] at hm + lift m to ℕ using hm.le + rw [Int.natCast_pos] at hm + simp only [norm_pow, zpow_neg, zpow_natCast, ← inv_pow] + exact (tendsto_pow_atTop hm.ne').comp NormedField.tendsto_norm_inverse_nhdsWithin_0_atTop + end NormedDivisionRing namespace NormedField @@ -361,6 +374,22 @@ theorem denseRange_nnnorm : DenseRange (nnnorm : α → ℝ≥0) := end Densely +section NontriviallyNormedField +variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {n : ℤ} {x : 𝕜} + +@[simp] +protected lemma continuousAt_zpow : ContinuousAt (fun x ↦ x ^ n) x ↔ x ≠ 0 ∨ 0 ≤ n := by + refine ⟨?_, continuousAt_zpow₀ _ _⟩ + contrapose! + rintro ⟨rfl, hm⟩ hc + exact not_tendsto_atTop_of_tendsto_nhds (hc.tendsto.mono_left nhdsWithin_le_nhds).norm + (tendsto_norm_zpow_nhdsWithin_0_atTop hm) + +@[simp] +protected lemma continuousAt_inv : ContinuousAt Inv.inv x ↔ x ≠ 0 := by + simpa using NormedField.continuousAt_zpow (n := -1) (x := x) + +end NontriviallyNormedField end NormedField namespace NNReal diff --git a/Mathlib/Analysis/Normed/Group/Basic.lean b/Mathlib/Analysis/Normed/Group/Basic.lean index 9ff5a866463b2..31e3b6bff3fe5 100644 --- a/Mathlib/Analysis/Normed/Group/Basic.lean +++ b/Mathlib/Analysis/Normed/Group/Basic.lean @@ -773,7 +773,8 @@ theorem tendsto_norm_div_self (x : E) : Tendsto (fun a => ‖a / x‖) (𝓝 x) theorem tendsto_norm' {x : E} : Tendsto (fun a => ‖a‖) (𝓝 x) (𝓝 ‖x‖) := by simpa using tendsto_id.dist (tendsto_const_nhds : Tendsto (fun _a => (1 : E)) _ _) -@[to_additive] +/-- See `tendsto_norm_one` for a version with pointed neighborhoods. -/ +@[to_additive "See `tendsto_norm_zero` for a version with pointed neighborhoods."] theorem tendsto_norm_one : Tendsto (fun a : E => ‖a‖) (𝓝 1) (𝓝 0) := by simpa using tendsto_norm_div_self (1 : E) @@ -1283,6 +1284,11 @@ theorem nnnorm_ne_zero_iff' : ‖a‖₊ ≠ 0 ↔ a ≠ 1 := @[to_additive (attr := simp) nnnorm_pos] lemma nnnorm_pos' : 0 < ‖a‖₊ ↔ a ≠ 1 := pos_iff_ne_zero.trans nnnorm_ne_zero_iff' +/-- See `tendsto_norm_one` for a version with full neighborhoods. -/ +@[to_additive "See `tendsto_norm_zero` for a version with full neighborhoods."] +lemma tendsto_norm_one' : Tendsto (norm : E → ℝ) (𝓝[≠] 1) (𝓝[>] 0) := + tendsto_norm_one.inf <| tendsto_principal_principal.2 fun _ hx ↦ norm_pos_iff''.2 hx + @[to_additive] theorem tendsto_norm_div_self_punctured_nhds (a : E) : Tendsto (fun x => ‖x / a‖) (𝓝[≠] a) (𝓝[>] 0) := @@ -1414,3 +1420,5 @@ instance (priority := 75) normedCommGroup [NormedCommGroup E] {S : Type*} [SetLi NormedCommGroup.induced _ _ (SubgroupClass.subtype s) Subtype.coe_injective end SubgroupClass + +lemma tendsto_norm_atTop_atTop : Tendsto (norm : ℝ → ℝ) atTop atTop := tendsto_abs_atTop_atTop diff --git a/Mathlib/Analysis/SpecificLimits/Normed.lean b/Mathlib/Analysis/SpecificLimits/Normed.lean index ba7281ba264e2..0e57727abb201 100644 --- a/Mathlib/Analysis/SpecificLimits/Normed.lean +++ b/Mathlib/Analysis/SpecificLimits/Normed.lean @@ -28,60 +28,8 @@ open Set Function Filter Finset Metric Asymptotics Topology Nat NNReal ENNReal variable {α β ι : Type*} -theorem tendsto_norm_atTop_atTop : Tendsto (norm : ℝ → ℝ) atTop atTop := - tendsto_abs_atTop_atTop - -theorem summable_of_absolute_convergence_real {f : ℕ → ℝ} : - (∃ r, Tendsto (fun n ↦ ∑ i ∈ range n, |f i|) atTop (𝓝 r)) → Summable f - | ⟨r, hr⟩ => by - refine .of_norm ⟨r, (hasSum_iff_tendsto_nat_of_nonneg ?_ _).2 ?_⟩ - · exact fun i ↦ norm_nonneg _ - · simpa only using hr - /-! ### Powers -/ - -theorem tendsto_norm_zero' {𝕜 : Type*} [NormedAddCommGroup 𝕜] : - Tendsto (norm : 𝕜 → ℝ) (𝓝[≠] 0) (𝓝[>] 0) := - tendsto_norm_zero.inf <| tendsto_principal_principal.2 fun _ hx ↦ norm_pos_iff.2 hx - -namespace NormedField - -theorem tendsto_norm_inverse_nhdsWithin_0_atTop {𝕜 : Type*} [NormedDivisionRing 𝕜] : - Tendsto (fun x : 𝕜 ↦ ‖x⁻¹‖) (𝓝[≠] 0) atTop := - (tendsto_inv_zero_atTop.comp tendsto_norm_zero').congr fun x ↦ (norm_inv x).symm - -theorem tendsto_norm_zpow_nhdsWithin_0_atTop {𝕜 : Type*} [NormedDivisionRing 𝕜] {m : ℤ} - (hm : m < 0) : - Tendsto (fun x : 𝕜 ↦ ‖x ^ m‖) (𝓝[≠] 0) atTop := by - rcases neg_surjective m with ⟨m, rfl⟩ - rw [neg_lt_zero] at hm; lift m to ℕ using hm.le; rw [Int.natCast_pos] at hm - simp only [norm_pow, zpow_neg, zpow_natCast, ← inv_pow] - exact (tendsto_pow_atTop hm.ne').comp NormedField.tendsto_norm_inverse_nhdsWithin_0_atTop - -/-- The (scalar) product of a sequence that tends to zero with a bounded one also tends to zero. -/ -theorem tendsto_zero_smul_of_tendsto_zero_of_bounded {ι 𝕜 𝔸 : Type*} [NormedDivisionRing 𝕜] - [NormedAddCommGroup 𝔸] [Module 𝕜 𝔸] [BoundedSMul 𝕜 𝔸] {l : Filter ι} {ε : ι → 𝕜} {f : ι → 𝔸} - (hε : Tendsto ε l (𝓝 0)) (hf : Filter.IsBoundedUnder (· ≤ ·) l (norm ∘ f)) : - Tendsto (ε • f) l (𝓝 0) := by - rw [← isLittleO_one_iff 𝕜] at hε ⊢ - simpa using IsLittleO.smul_isBigO hε (hf.isBigO_const (one_ne_zero : (1 : 𝕜) ≠ 0)) - -@[simp] -theorem continuousAt_zpow {𝕜 : Type*} [NontriviallyNormedField 𝕜] {m : ℤ} {x : 𝕜} : - ContinuousAt (fun x ↦ x ^ m) x ↔ x ≠ 0 ∨ 0 ≤ m := by - refine ⟨?_, continuousAt_zpow₀ _ _⟩ - contrapose!; rintro ⟨rfl, hm⟩ hc - exact not_tendsto_atTop_of_tendsto_nhds (hc.tendsto.mono_left nhdsWithin_le_nhds).norm - (tendsto_norm_zpow_nhdsWithin_0_atTop hm) - -@[simp] -theorem continuousAt_inv {𝕜 : Type*} [NontriviallyNormedField 𝕜] {x : 𝕜} : - ContinuousAt Inv.inv x ↔ x ≠ 0 := by - simpa [(zero_lt_one' ℤ).not_le] using @continuousAt_zpow _ _ (-1) x - -end NormedField - theorem isLittleO_pow_pow_of_lt_left {r₁ r₂ : ℝ} (h₁ : 0 ≤ r₁) (h₂ : r₁ < r₂) : (fun n : ℕ ↦ r₁ ^ n) =o[atTop] fun n ↦ r₂ ^ n := have H : 0 < r₂ := h₁.trans_lt h₂ From e5f5585586d47ff9fbe14e98764da0efad352e2a Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Tue, 22 Oct 2024 11:41:54 +0000 Subject: [PATCH 268/505] feat(Order/ConditionallyCompleteLattice): `lt_ciSup_iff` and variants (#17965) --- .../ConditionallyCompleteLattice/Basic.lean | 30 ++++++++++++++++++- Mathlib/SetTheory/Ordinal/Arithmetic.lean | 11 ++++--- 2 files changed, 34 insertions(+), 7 deletions(-) diff --git a/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean b/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean index e0b7205db6894..614bad616a8b8 100644 --- a/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean +++ b/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean @@ -997,6 +997,20 @@ When `iInf f < a`, there is an element `i` such that `f i < a`. theorem exists_lt_of_ciInf_lt [Nonempty ι] {f : ι → α} (h : iInf f < a) : ∃ i, f i < a := exists_lt_of_lt_ciSup (α := αᵒᵈ) h +theorem lt_csSup_iff (hb : BddAbove s) (hs : s.Nonempty) : a < sSup s ↔ ∃ b ∈ s, a < b := by + simpa only [not_le, not_forall₂, exists_prop] using (csSup_le_iff hb hs (a := a)).not + +theorem csInf_lt_iff (hb : BddBelow s) (hs : s.Nonempty) : sInf s < a ↔ ∃ b ∈ s, b < a := by + simpa only [not_le, not_forall₂, exists_prop] using (le_csInf_iff hb hs).not + +theorem lt_ciSup_iff [Nonempty ι] {f : ι → α} (hb : BddAbove (range f)) : + a < iSup f ↔ ∃ i, a < f i := by + simpa only [mem_range, exists_exists_eq_and] using lt_csSup_iff hb (range_nonempty _) + +theorem ciInf_lt_iff [Nonempty ι] {f : ι → α} (hb : BddBelow (range f)) : + iInf f < a ↔ ∃ i, f i < a := by + simpa only [mem_range, exists_exists_eq_and] using csInf_lt_iff hb (range_nonempty _) + theorem csSup_of_not_bddAbove {s : Set α} (hs : ¬BddAbove s) : sSup s = sSup ∅ := ConditionallyCompleteLinearOrder.csSup_of_not_bddAbove s hs @@ -1179,12 +1193,19 @@ theorem isLUB_csSup' {s : Set α} (hs : BddAbove s) : IsLUB s (sSup s) := by · simp only [csSup_empty, isLUB_empty] · exact isLUB_csSup hne hs +/-- In conditionally complete orders with a bottom element, the nonempty condition can be omitted +from `csSup_le_iff`. -/ theorem csSup_le_iff' {s : Set α} (hs : BddAbove s) {a : α} : sSup s ≤ a ↔ ∀ x ∈ s, x ≤ a := isLUB_le_iff (isLUB_csSup' hs) theorem csSup_le' {s : Set α} {a : α} (h : a ∈ upperBounds s) : sSup s ≤ a := (csSup_le_iff' ⟨a, h⟩).2 h +/-- In conditionally complete orders with a bottom element, the nonempty condition can be omitted +from `lt_csSup_iff`. -/ +theorem lt_csSup_iff' (hb : BddAbove s) : a < sSup s ↔ ∃ b ∈ s, a < b := by + simpa only [not_le, not_forall₂, exists_prop] using (csSup_le_iff' hb).not + theorem le_csSup_iff' {s : Set α} {a : α} (h : BddAbove s) : a ≤ sSup s ↔ ∀ b, b ∈ upperBounds s → a ≤ b := ⟨fun h _ hb => le_trans h (csSup_le' hb), fun hb => hb _ fun _ => le_csSup h⟩ @@ -1209,6 +1230,8 @@ theorem exists_lt_of_lt_csSup' {s : Set α} {a : α} (h : a < sSup s) : ∃ b contrapose! h exact csSup_le' h +/-- In conditionally complete orders with a bottom element, the nonempty condition can be omitted +from `ciSup_le_iff`. -/ theorem ciSup_le_iff' {f : ι → α} (h : BddAbove (range f)) {a : α} : ⨆ i, f i ≤ a ↔ ∀ i, f i ≤ a := (csSup_le_iff' h).trans forall_mem_range @@ -1216,6 +1239,11 @@ theorem ciSup_le_iff' {f : ι → α} (h : BddAbove (range f)) {a : α} : theorem ciSup_le' {f : ι → α} {a : α} (h : ∀ i, f i ≤ a) : ⨆ i, f i ≤ a := csSup_le' <| forall_mem_range.2 h +/-- In conditionally complete orders with a bottom element, the nonempty condition can be omitted +from `lt_ciSup_iff`. -/ +theorem lt_ciSup_iff' {f : ι → α} (h : BddAbove (range f)) : a < iSup f ↔ ∃ i, a < f i := by + simpa only [not_le, not_forall] using (ciSup_le_iff' h).not + theorem exists_lt_of_lt_ciSup' {f : ι → α} {a : α} (h : a < ⨆ i, f i) : ∃ i, a < f i := by contrapose! h exact ciSup_le' h @@ -1663,4 +1691,4 @@ lemma iInf_coe_lt_top : ⨅ i, (f i : WithTop α) < ⊤ ↔ Nonempty ι := by end WithTop end WithTopBot -set_option linter.style.longFile 1700 +set_option linter.style.longFile 1800 diff --git a/Mathlib/SetTheory/Ordinal/Arithmetic.lean b/Mathlib/SetTheory/Ordinal/Arithmetic.lean index c8a10d44c7ce0..3035879d9e2b4 100644 --- a/Mathlib/SetTheory/Ordinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Ordinal/Arithmetic.lean @@ -1204,8 +1204,8 @@ set_option linter.deprecated false in theorem sup_le_iff {ι : Type u} {f : ι → Ordinal.{max u v}} {a} : sup.{_, v} f ≤ a ↔ ∀ i, f i ≤ a := Ordinal.iSup_le_iff -/-- `ciSup_le'` whenever the input type is small in the output universe. -/ -protected theorem iSup_le {ι} {f : ι → Ordinal.{u}} {a} : +/-- An alias of `ciSup_le'` for discoverability. -/ +protected theorem iSup_le {ι} {f : ι → Ordinal} {a} : (∀ i, f i ≤ a) → iSup f ≤ a := ciSup_le' @@ -1214,11 +1214,10 @@ set_option linter.deprecated false in theorem sup_le {ι : Type u} {f : ι → Ordinal.{max u v}} {a} : (∀ i, f i ≤ a) → sup.{_, v} f ≤ a := Ordinal.iSup_le --- TODO: generalize to conditionally complete linear orders. +/-- `lt_ciSup_iff'` whenever the input type is small in the output universe. -/ protected theorem lt_iSup {ι} {f : ι → Ordinal.{u}} {a : Ordinal.{u}} [Small.{u} ι] : - a < iSup f ↔ ∃ i, a < f i := by - rw [← not_iff_not] - simpa using Ordinal.iSup_le_iff + a < iSup f ↔ ∃ i, a < f i := + lt_ciSup_iff' (bddAbove_of_small _) set_option linter.deprecated false in @[deprecated Ordinal.lt_iSup (since := "2024-08-27")] From 9e395964efc61b06def256422ee14a1ad8734a2a Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 22 Oct 2024 12:13:54 +0000 Subject: [PATCH 269/505] chore(Archive): use newly introduced finset notation (#18052) --- Archive/Imo/Imo1994Q1.lean | 2 +- Archive/Imo/Imo2021Q1.lean | 4 +- Archive/Sensitivity.lean | 23 ++++--- .../AscendingDescendingSequences.lean | 4 +- Archive/Wiedijk100Theorems/Partition.lean | 38 ++++-------- .../SumOfPrimeReciprocalsDiverges.lean | 60 +++++++++---------- 6 files changed, 56 insertions(+), 75 deletions(-) diff --git a/Archive/Imo/Imo1994Q1.lean b/Archive/Imo/Imo1994Q1.lean index 6ea9e04e8862f..4420507ec7f59 100644 --- a/Archive/Imo/Imo1994Q1.lean +++ b/Archive/Imo/Imo1994Q1.lean @@ -43,7 +43,7 @@ end Imo1994Q1 open Imo1994Q1 -theorem imo1994_q1 (n : ℕ) (m : ℕ) (A : Finset ℕ) (hm : A.card = m + 1) +theorem imo1994_q1 (n : ℕ) (m : ℕ) (A : Finset ℕ) (hm : #A = m + 1) (hrange : ∀ a ∈ A, 0 < a ∧ a ≤ n) (hadd : ∀ a ∈ A, ∀ b ∈ A, a + b ≤ n → a + b ∈ A) : (m + 1) * (n + 1) ≤ 2 * ∑ x ∈ A, x := by diff --git a/Archive/Imo/Imo2021Q1.lean b/Archive/Imo/Imo2021Q1.lean index 051bea2e28ac9..54f3cdf9ec531 100644 --- a/Archive/Imo/Imo2021Q1.lean +++ b/Archive/Imo/Imo2021Q1.lean @@ -76,7 +76,7 @@ lemma exists_triplet_summing_to_squares {n : ℕ} (hn : 100 ≤ n) : -- pair of pairwise unequal elements of B sums to a perfect square. lemma exists_finset_3_le_card_with_pairs_summing_to_squares {n : ℕ} (hn : 100 ≤ n) : ∃ B : Finset ℕ, - 2 * 1 + 1 ≤ B.card ∧ + 2 * 1 + 1 ≤ #B ∧ (∀ a ∈ B, ∀ b ∈ B, a ≠ b → IsSquare (a + b)) ∧ ∀ c ∈ B, n ≤ c ∧ c ≤ 2 * n := by obtain ⟨a, b, c, hna, hab, hbc, hcn, h₁, h₂, h₃⟩ := exists_triplet_summing_to_squares hn @@ -115,7 +115,7 @@ theorem imo2021_q1 : obtain ⟨B, hB, h₁, h₂⟩ := exists_finset_3_le_card_with_pairs_summing_to_squares hn have hBsub : B ⊆ Finset.Icc n (2 * n) := by intro c hcB; simpa only [Finset.mem_Icc] using h₂ c hcB - have hB' : 2 * 1 < (B ∩ (Finset.Icc n (2 * n) \ A) ∪ B ∩ A).card := by + have hB' : 2 * 1 < #(B ∩ (Icc n (2 * n) \ A) ∪ B ∩ A) := by rwa [← inter_union_distrib_left, sdiff_union_self_eq_union, union_eq_left.2 hA, inter_eq_left.2 hBsub, ← Nat.succ_le_iff] -- Since B has cardinality greater or equal to 3, there must exist a subset C ⊆ B such that diff --git a/Archive/Sensitivity.lean b/Archive/Sensitivity.lean index 5cae8d062537c..ab52f935c6633 100644 --- a/Archive/Sensitivity.lean +++ b/Archive/Sensitivity.lean @@ -41,7 +41,7 @@ noncomputable section local notation "√" => Real.sqrt -open Function Bool LinearMap Fintype Module Module.DualBases +open Bool Finset Fintype Function LinearMap Module Module.DualBases /-! ### The hypercube @@ -357,7 +357,7 @@ local notation "Span" => Submodule.span ℝ natural number. -/ -local notation "Card " X:70 => Finset.card (Set.toFinset X) +local notation "Card " X:70 => #(Set.toFinset X) /-! In the following, `⊓` and `⊔` will denote intersection and sums of ℝ-subspaces, equipped with their subspace structures. The notations come from the general @@ -431,20 +431,19 @@ theorem huang_degree_theorem (H : Set (Q m.succ)) (hH : Card H ≥ 2 ^ m + 1) : (norm_sum_le _ fun p => coeffs y p * _) _ = ∑ p ∈ (coeffs y).support, |coeffs y p| * ite (p ∈ q.adjacent) 1 0 := by simp only [abs_mul, f_matrix] - _ = ∑ p ∈ (coeffs y).support.filter q.adjacent, |coeffs y p| := by - simp [Finset.sum_filter]; rfl - _ ≤ ∑ _p ∈ (coeffs y).support.filter q.adjacent, |coeffs y q| := - (Finset.sum_le_sum fun p _ => H_max p) - _ = (((coeffs y).support.filter q.adjacent).card : ℝ) * |coeffs y q| := by - rw [Finset.sum_const, nsmul_eq_mul] - _ = (((coeffs y).support ∩ q.adjacent.toFinset).card : ℝ) * |coeffs y q| := by + _ = ∑ p ∈ (coeffs y).support with q.adjacent p, |coeffs y p| := by + simp [sum_filter]; rfl + _ ≤ ∑ p ∈ (coeffs y).support with q.adjacent p, |coeffs y q| := sum_le_sum fun p _ ↦ H_max p + _ = #{p ∈ (coeffs y).support | q.adjacent p} * |coeffs y q| := by + rw [sum_const, nsmul_eq_mul] + _ = #((coeffs y).support ∩ q.adjacent.toFinset) * |coeffs y q| := by congr with x; simp; rfl - _ ≤ Finset.card (H ∩ q.adjacent).toFinset * |ε q y| := by + _ ≤ #(H ∩ q.adjacent).toFinset * |ε q y| := by refine (mul_le_mul_right H_q_pos).2 ?_ norm_cast - apply Finset.card_le_card + apply card_le_card rw [Set.toFinset_inter] - convert Finset.inter_subset_inter_right coeffs_support + convert inter_subset_inter_right coeffs_support end diff --git a/Archive/Wiedijk100Theorems/AscendingDescendingSequences.lean b/Archive/Wiedijk100Theorems/AscendingDescendingSequences.lean index df05d9256c5a5..884fb11207295 100644 --- a/Archive/Wiedijk100Theorems/AscendingDescendingSequences.lean +++ b/Archive/Wiedijk100Theorems/AscendingDescendingSequences.lean @@ -39,8 +39,8 @@ We then show the pair of labels must be unique. Now if there is no increasing se which is a contradiction if there are more than `r * s` elements. -/ theorem erdos_szekeres {r s n : ℕ} {f : Fin n → α} (hn : r * s < n) (hf : Injective f) : - (∃ t : Finset (Fin n), r < t.card ∧ StrictMonoOn f ↑t) ∨ - ∃ t : Finset (Fin n), s < t.card ∧ StrictAntiOn f ↑t := by + (∃ t : Finset (Fin n), r < #t ∧ StrictMonoOn f ↑t) ∨ + ∃ t : Finset (Fin n), s < #t ∧ StrictAntiOn f ↑t := by -- Given an index `i`, produce the set of increasing (resp., decreasing) subsequences which ends -- at `i`. let inc_sequences_ending_in : Fin n → Finset (Finset (Fin n)) := fun i => diff --git a/Archive/Wiedijk100Theorems/Partition.lean b/Archive/Wiedijk100Theorems/Partition.lean index 689d098bcd228..96f1877c186a5 100644 --- a/Archive/Wiedijk100Theorems/Partition.lean +++ b/Archive/Wiedijk100Theorems/Partition.lean @@ -87,7 +87,7 @@ or `ℝ`. def partialDistinctGF (m : ℕ) [CommSemiring α] := ∏ i ∈ range m, (1 + (X : PowerSeries α) ^ (i + 1)) -open Finset.HasAntidiagonal Finset +open Finset.HasAntidiagonal universe u variable {ι : Type u} @@ -134,9 +134,7 @@ theorem num_series' [Field α] (i : ℕ) : sub_eq_iff_eq_add, zero_add] symm split_ifs with h - · suffices - ((antidiagonal (n+1)).filter fun a : ℕ × ℕ => i + 1 ∣ a.fst ∧ a.snd = i + 1).card = - 1 by + · suffices #{a ∈ antidiagonal (n + 1) | i + 1 ∣ a.fst ∧ a.snd = i + 1} = 1 by simp only [Set.mem_setOf_eq]; convert congr_arg ((↑) : ℕ → α) this; norm_cast rw [card_eq_one] cases' h with p hp @@ -151,9 +149,7 @@ theorem num_series' [Field α] (i : ℕ) : match p with | 0 => rw [mul_zero] at hp; cases hp | p + 1 => rw [hp]; simp [mul_add] - · suffices - (filter (fun a : ℕ × ℕ => i + 1 ∣ a.fst ∧ a.snd = i + 1) (antidiagonal (n+1))).card = - 0 by + · suffices #{a ∈ antidiagonal (n + 1) | i + 1 ∣ a.fst ∧ a.snd = i + 1} = 0 by simp only [Set.mem_setOf_eq]; convert congr_arg ((↑) : ℕ → α) this; norm_cast rw [card_eq_zero] apply eq_empty_of_forall_not_mem @@ -169,11 +165,8 @@ def mkOdd : ℕ ↪ ℕ := -- The main workhorse of the partition theorem proof. theorem partialGF_prop (α : Type*) [CommSemiring α] (n : ℕ) (s : Finset ℕ) (hs : ∀ i ∈ s, 0 < i) (c : ℕ → Set ℕ) (hc : ∀ i, i ∉ s → 0 ∈ c i) : - (Finset.card - ((univ : Finset (Nat.Partition n)).filter fun p => - (∀ j, p.parts.count j ∈ c j) ∧ ∀ j ∈ p.parts, j ∈ s) : - α) = - (coeff α n) (∏ i ∈ s, indicatorSeries α ((· * i) '' c i)) := by + #{p : n.Partition | (∀ j, p.parts.count j ∈ c j) ∧ ∀ j ∈ p.parts, j ∈ s} = + coeff α n (∏ i ∈ s, indicatorSeries α ((· * i) '' c i)) := by simp_rw [coeff_prod, coeff_indicator, prod_boole, sum_boole] apply congr_arg simp only [mem_univ, forall_true_left, not_and, not_forall, exists_prop, @@ -257,11 +250,7 @@ theorem partialGF_prop (α : Type*) [CommSemiring α] (n : ℕ) (s : Finset ℕ) exact not_mem_mono hf'.2 h theorem partialOddGF_prop [Field α] (n m : ℕ) : - (Finset.card - ((univ : Finset (Nat.Partition n)).filter fun p => - ∀ j ∈ p.parts, j ∈ (range m).map mkOdd) : - α) = - coeff α n (partialOddGF m) := by + #{p : n.Partition | ∀ j ∈ p.parts, j ∈ (range m).map mkOdd} = coeff α n (partialOddGF m) := by rw [partialOddGF] convert partialGF_prop α n ((range m).map mkOdd) _ (fun _ => Set.univ) (fun _ _ => trivial) using 2 @@ -284,7 +273,7 @@ theorem partialOddGF_prop [Field α] (n m : ℕ) : /-- If m is big enough, the partial product's coefficient counts the number of odd partitions -/ theorem oddGF_prop [Field α] (n m : ℕ) (h : n < m * 2) : - (Finset.card (Nat.Partition.odds n) : α) = coeff α n (partialOddGF m) := by + #(Nat.Partition.odds n) = coeff α n (partialOddGF m) := by rw [← partialOddGF_prop, Nat.Partition.odds] congr with p apply forall₂_congr @@ -305,10 +294,8 @@ theorem oddGF_prop [Field α] (n m : ℕ) (h : n < m * 2) : apply Nat.two_not_dvd_two_mul_add_one theorem partialDistinctGF_prop [CommSemiring α] (n m : ℕ) : - (Finset.card - ((univ : Finset (Nat.Partition n)).filter fun p => - p.parts.Nodup ∧ ∀ j ∈ p.parts, j ∈ (range m).map ⟨Nat.succ, Nat.succ_injective⟩) : - α) = + #{p : n.Partition | + p.parts.Nodup ∧ ∀ j ∈ p.parts, j ∈ (range m).map ⟨Nat.succ, Nat.succ_injective⟩} = coeff α n (partialDistinctGF m) := by rw [partialDistinctGF] convert partialGF_prop α n @@ -328,7 +315,7 @@ theorem partialDistinctGF_prop [CommSemiring α] (n m : ℕ) : /-- If m is big enough, the partial product's coefficient counts the number of distinct partitions -/ theorem distinctGF_prop [CommSemiring α] (n m : ℕ) (h : n < m + 1) : - ((Nat.Partition.distincts n).card : α) = coeff α n (partialDistinctGF m) := by + #(Nat.Partition.distincts n) = coeff α n (partialDistinctGF m) := by rw [← partialDistinctGF_prop, Nat.Partition.distincts] congr with p apply (and_iff_left _).symm @@ -388,10 +375,9 @@ theorem same_coeffs [Field α] (m n : ℕ) (h : n ≤ m) : rw [order_X_pow] exact mod_cast Nat.lt_succ_of_le (le_add_right h) -theorem partition_theorem (n : ℕ) : - (Nat.Partition.odds n).card = (Nat.Partition.distincts n).card := by +theorem partition_theorem (n : ℕ) : #(Nat.Partition.odds n) = #(Nat.Partition.distincts n) := by -- We need the counts to live in some field (which contains ℕ), so let's just use ℚ - suffices ((Nat.Partition.odds n).card : ℚ) = (Nat.Partition.distincts n).card from + suffices (#(Nat.Partition.odds n) : ℚ) = #(Nat.Partition.distincts n) from mod_cast this rw [distinctGF_prop n (n + 1) (by linarith)] rw [oddGF_prop n (n + 1) (by linarith)] diff --git a/Archive/Wiedijk100Theorems/SumOfPrimeReciprocalsDiverges.lean b/Archive/Wiedijk100Theorems/SumOfPrimeReciprocalsDiverges.lean index e30881c10059f..914eeefba8a81 100644 --- a/Archive/Wiedijk100Theorems/SumOfPrimeReciprocalsDiverges.lean +++ b/Archive/Wiedijk100Theorems/SumOfPrimeReciprocalsDiverges.lean @@ -44,20 +44,17 @@ namespace Theorems100 /-- The primes in `(k, x]`. -/ -def P (x k : ℕ) : Finset ℕ := - Finset.filter (fun p => k < p ∧ Nat.Prime p) (range (x + 1)) +def P (x k : ℕ) : Finset ℕ := {p ∈ range (x + 1) | k < p ∧ p.Prime} /-- The union over those primes `p ∈ (k, x]` of the sets of `e < x` for which `e + 1` is a multiple of `p`, i.e., those `e < x` for which there is a prime `p ∈ (k, x]` that divides `e + 1`. -/ -def U (x k : ℕ) := - Finset.biUnion (P x k) (fun p => Finset.filter (fun e => p ∣ e + 1) (range x)) +def U (x k : ℕ) : Finset ℕ := (P x k).biUnion fun p ↦ {e ∈ range x | p ∣ e + 1} open Classical in /-- Those `e < x` for which `e + 1` is a product of powers of primes smaller than or equal to `k`. -/ -noncomputable def M (x k : ℕ) := - Finset.filter (fun e => ∀ p : ℕ, Nat.Prime p ∧ p ∣ e + 1 → p ≤ k) (range x) +noncomputable def M (x k : ℕ) : Finset ℕ := {e ∈ range x | ∀ p : ℕ, p.Prime ∧ p ∣ e + 1 → p ≤ k} /-- If the sum of the reciprocals of the primes converges, there exists a `k : ℕ` such that the sum of @@ -67,11 +64,11 @@ More precisely, for any `x : ℕ`, the sum of the reciprocals of the primes betw is less than 1/2. -/ theorem sum_lt_half_of_not_tendsto - (h : ¬Tendsto (fun n => ∑ p ∈ Finset.filter (fun p => Nat.Prime p) (range n), 1 / (p : ℝ)) + (h : ¬Tendsto (fun n => ∑ p ∈ range n with p.Prime, 1 / (p : ℝ)) atTop atTop) : ∃ k, ∀ x, ∑ p ∈ P x k, 1 / (p : ℝ) < 1 / 2 := by have h0 : - (fun n => ∑ p ∈ Finset.filter (fun p => Nat.Prime p) (range n), 1 / (p : ℝ)) = fun n => + (fun n => ∑ p ∈ range n with p.Prime, 1 / (p : ℝ)) = fun n => ∑ p ∈ range n, ite (Nat.Prime p) (1 / (p : ℝ)) 0 := by simp only [sum_filter] have hf : ∀ n : ℕ, 0 ≤ ite (Nat.Prime n) (1 / (n : ℝ)) 0 := by @@ -113,12 +110,12 @@ theorem range_sdiff_eq_biUnion {x k : ℕ} : range x \ M x k = U x k := by The number of `e < x` for which `e + 1` has a prime factor `p > k` is bounded by `x` times the sum of reciprocals of primes in `(k, x]`. -/ -theorem card_le_mul_sum {x k : ℕ} : (card (U x k) : ℝ) ≤ x * ∑ p ∈ P x k, 1 / (p : ℝ) := by - let P := Finset.filter (fun p => k < p ∧ Nat.Prime p) (range (x + 1)) - let N p := Finset.filter (fun e => p ∣ e + 1) (range x) - have h : card (Finset.biUnion P N) ≤ ∑ p ∈ P, card (N p) := card_biUnion_le +theorem card_le_mul_sum {x k : ℕ} : #(U x k) ≤ x * ∑ p ∈ P x k, 1 / (p : ℝ) := by + let P := {p ∈ range (x + 1) | k < p ∧ p.Prime} + let N p := {e ∈ range x | p ∣ e + 1} + have h : #(P.biUnion N) ≤ ∑ p ∈ P, #(N p) := card_biUnion_le calc - (card (Finset.biUnion P N) : ℝ) ≤ ∑ p ∈ P, (card (N p) : ℝ) := by assumption_mod_cast + (#(P.biUnion N) : ℝ) ≤ ∑ p ∈ P, (#(N p) : ℝ) := by assumption_mod_cast _ ≤ ∑ p ∈ P, x * (1 / (p : ℝ)) := sum_le_sum fun p _ => ?_ _ = x * ∑ p ∈ P, 1 / (p : ℝ) := by rw [mul_sum] simp only [N, mul_one_div, Nat.card_multiples, Nat.cast_div_le] @@ -127,9 +124,8 @@ theorem card_le_mul_sum {x k : ℕ} : (card (U x k) : ℝ) ≤ x * ∑ p ∈ P x The number of `e < x` for which `e + 1` is a squarefree product of primes smaller than or equal to `k` is bounded by `2 ^ k`, the number of subsets of `[1, k]`. -/ -theorem card_le_two_pow {x k : ℕ} : - card (Finset.filter (fun e => Squarefree (e + 1)) (M x k)) ≤ 2 ^ k := by - let M₁ := Finset.filter (fun e => Squarefree (e + 1)) (M x k) +theorem card_le_two_pow {x k : ℕ} : #{e ∈ M x k | Squarefree (e + 1)} ≤ 2 ^ k := by + let M₁ := {e ∈ M x k | Squarefree (e + 1)} let f s := (∏ a ∈ s, a) - 1 let K := powerset (image Nat.succ (range k)) -- Take `e` in `M x k`. If `e + 1` is squarefree, then it is the product of a subset of `[1, k]`. @@ -151,18 +147,18 @@ theorem card_le_two_pow {x k : ℕ} : -- The number of elements of `M x k` with `e + 1` squarefree is bounded by the number of subsets -- of `[1, k]`. calc - 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_right_mono₀ one_le_two card_image_le + #M₁ ≤ #(image f K) := card_le_card h + _ ≤ #K := card_image_le + _ ≤ 2 ^ #(image Nat.succ (range k)) := by simp only [K, card_powerset]; rfl + _ ≤ 2 ^ #(range k) := pow_right_mono₀ one_le_two card_image_le _ = 2 ^ k := by rw [card_range k] /-- The number of `e < x` for which `e + 1` is a product of powers of primes smaller than or equal to `k` is bounded by `2 ^ k * nat.sqrt x`. -/ -theorem card_le_two_pow_mul_sqrt {x k : ℕ} : card (M x k) ≤ 2 ^ k * Nat.sqrt x := by - let M₁ := Finset.filter (fun e => Squarefree (e + 1)) (M x k) +theorem card_le_two_pow_mul_sqrt {x k : ℕ} : #(M x k) ≤ 2 ^ k * Nat.sqrt x := by + let M₁ := {e ∈ M x k | Squarefree (e + 1)} let M₂ := M (Nat.sqrt x) k let K := M₁ ×ˢ M₂ let f : ℕ × ℕ → ℕ := fun mn => (mn.2 + 1) ^ 2 * (mn.1 + 1) - 1 @@ -185,16 +181,16 @@ theorem card_le_two_pow_mul_sqrt {x k : ℕ} : card (M x k) ≤ 2 ^ k * Nat.sqrt _ ≤ (m + 1).sqrt := by simpa only [Nat.le_sqrt, pow_two] using Nat.le_of_dvd hm' hbm _ ≤ x.sqrt := Nat.sqrt_le_sqrt (Nat.succ_le_iff.mpr hm.1) · exact hm.2 p ⟨hp.1, hp.2.trans (Nat.dvd_of_pow_dvd one_le_two hbm)⟩ - have h2 : card M₂ ≤ Nat.sqrt x := by + have h2 : #M₂ ≤ Nat.sqrt x := by rw [← card_range (Nat.sqrt x)]; apply card_le_card; simp [M, M₂] calc - card (M x k) ≤ card (image f K) := card_le_card h1 - _ ≤ card K := card_image_le - _ = card M₁ * card M₂ := card_product M₁ M₂ + #(M x k) ≤ #(image f K) := card_le_card h1 + _ ≤ #K := card_image_le + _ = #M₁ * #M₂ := card_product M₁ M₂ _ ≤ 2 ^ k * x.sqrt := mul_le_mul' card_le_two_pow h2 theorem Real.tendsto_sum_one_div_prime_atTop : - Tendsto (fun n => ∑ p ∈ Finset.filter (fun p => Nat.Prime p) (range n), 1 / (p : ℝ)) + Tendsto (fun n => ∑ p ∈ range n with p.Prime, 1 / (p : ℝ)) atTop atTop := by -- Assume that the sum of the reciprocals of the primes converges. by_contra h @@ -209,10 +205,10 @@ theorem Real.tendsto_sum_one_div_prime_atTop : -- than or equal to `k`; set M' := M x k with hM' -- * `U`, the subset of those `e` for which there is a prime `p > k` that divides `e + 1`. - let P := Finset.filter (fun p => k < p ∧ Nat.Prime p) (range (x + 1)) + let P := {p ∈ range (x + 1) | k < p ∧ p.Prime} set U' := U x k with hU' -- This is indeed a partition, so `|U| + |M| = |range x| = x`. - have h2 : x = card U' + card M' := by + have h2 : x = #U' + #M' := by rw [← card_range x, hU', hM', ← range_sdiff_eq_biUnion] classical exact (card_sdiff_add_card_eq_card (Finset.filter_subset _ _)).symm @@ -220,17 +216,17 @@ theorem Real.tendsto_sum_one_div_prime_atTop : -- and for U, the inequality is strict. have h3 := calc - (card U' : ℝ) ≤ x * ∑ p ∈ P, 1 / (p : ℝ) := card_le_mul_sum + (#U' : ℝ) ≤ x * ∑ p ∈ P, 1 / (p : ℝ) := card_le_mul_sum _ < x * (1 / 2) := mul_lt_mul_of_pos_left (h1 x) (by norm_num [x]) _ = x / 2 := mul_one_div (x : ℝ) 2 have h4 := calc - (card M' : ℝ) ≤ 2 ^ k * x.sqrt := by exact mod_cast card_le_two_pow_mul_sqrt + (#M' : ℝ) ≤ 2 ^ k * x.sqrt := by exact mod_cast card_le_two_pow_mul_sqrt _ = 2 ^ k * (2 ^ (k + 1) : ℕ) := by rw [Nat.sqrt_eq] _ = x / 2 := by field_simp [x, mul_right_comm, ← pow_succ] refine lt_irrefl (x : ℝ) ?_ calc - (x : ℝ) = (card U' : ℝ) + (card M' : ℝ) := by assumption_mod_cast + (x : ℝ) = (#U' : ℝ) + (#M' : ℝ) := by assumption_mod_cast _ < x / 2 + x / 2 := add_lt_add_of_lt_of_le h3 h4 _ = x := add_halves (x : ℝ) From 90daa70a1c21e87cf0f400bbd4e55639989845f1 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 22 Oct 2024 12:48:11 +0000 Subject: [PATCH 270/505] chore(Algebra): use newly introduced finset notation (#18053) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit For `s : Finset α`, replace `s.card` with `#s`, `s.filter p` with `{x ∈ s | p x}`, `∑ x ∈ s.filter p, f x` with `∑ x ∈ s with p x, f x`. Rewrap lines around and open the `Finset` locale where necessary. --- Mathlib/Algebra/BigOperators/Expect.lean | 22 +-- .../Algebra/BigOperators/Group/Finset.lean | 153 +++++++++--------- Mathlib/Algebra/BigOperators/Ring.lean | 24 ++- Mathlib/Algebra/GeomSum.lean | 2 +- Mathlib/Algebra/Group/UniqueProds/Basic.lean | 28 ++-- Mathlib/Algebra/Module/BigOperators.lean | 8 +- Mathlib/Algebra/MonoidAlgebra/Defs.lean | 10 +- .../Order/BigOperators/Group/Finset.lean | 78 ++++----- .../Algebra/Order/CauSeq/BigOperators.lean | 4 +- Mathlib/Algebra/Order/Chebyshev.lean | 28 ++-- Mathlib/Algebra/Polynomial/Basic.lean | 4 +- Mathlib/Algebra/Polynomial/BigOperators.lean | 6 +- Mathlib/Algebra/Polynomial/Coeff.lean | 14 +- .../Polynomial/Degree/Definitions.lean | 8 +- Mathlib/Algebra/Polynomial/EraseLead.lean | 35 ++-- Mathlib/Algebra/Polynomial/Reverse.lean | 4 +- Mathlib/Algebra/Polynomial/Roots.lean | 16 +- Mathlib/Algebra/Polynomial/UnitTrinomial.lean | 4 +- .../SimpleGraph/Regularity/Increment.lean | 2 +- 19 files changed, 215 insertions(+), 235 deletions(-) diff --git a/Mathlib/Algebra/BigOperators/Expect.lean b/Mathlib/Algebra/BigOperators/Expect.lean index 22969768a5f4f..df5560dfae390 100644 --- a/Mathlib/Algebra/BigOperators/Expect.lean +++ b/Mathlib/Algebra/BigOperators/Expect.lean @@ -53,7 +53,7 @@ local notation a " /ℚ " q => (q : ℚ≥0)⁻¹ • a /-- Average of a function over a finset. If the finset is empty, this is equal to zero. -/ def Finset.expect [AddCommMonoid M] [Module ℚ≥0 M] (s : Finset ι) (f : ι → M) : M := - (s.card : ℚ≥0)⁻¹ • ∑ i ∈ s, f i + (#s : ℚ≥0)⁻¹ • ∑ i ∈ s, f i namespace BigOperators open Batteries.ExtendedBinder Lean Meta @@ -152,36 +152,36 @@ lemma expect_add_expect_comm (f₁ f₂ g₁ g₂ : ι → M) : simp_rw [expect_add_distrib, add_add_add_comm] lemma expect_eq_single_of_mem (i : ι) (hi : i ∈ s) (h : ∀ j ∈ s, j ≠ i → f j = 0) : - 𝔼 i ∈ s, f i = f i /ℚ s.card := by rw [expect, sum_eq_single_of_mem _ hi h] + 𝔼 i ∈ s, f i = f i /ℚ #s := by rw [expect, sum_eq_single_of_mem _ hi h] lemma expect_ite_zero (s : Finset ι) (p : ι → Prop) [DecidablePred p] (h : ∀ i ∈ s, ∀ j ∈ s, p i → p j → i = j) (a : M) : - 𝔼 i ∈ s, ite (p i) a 0 = ite (∃ i ∈ s, p i) (a /ℚ s.card) 0 := by + 𝔼 i ∈ s, ite (p i) a 0 = ite (∃ i ∈ s, p i) (a /ℚ #s) 0 := by split_ifs <;> simp [expect, sum_ite_zero _ _ h, *] section DecidableEq variable [DecidableEq ι] lemma expect_ite_mem (s t : Finset ι) (f : ι → M) : - 𝔼 i ∈ s, (if i ∈ t then f i else 0) = ((s ∩ t).card / s.card : ℚ≥0) • 𝔼 i ∈ s ∩ t, f i := by + 𝔼 i ∈ s, (if i ∈ t then f i else 0) = (#(s ∩ t) / #s : ℚ≥0) • 𝔼 i ∈ s ∩ t, f i := by obtain hst | hst := (s ∩ t).eq_empty_or_nonempty · simp [expect, hst] · simp [expect, smul_smul, ← inv_mul_eq_div, hst.card_ne_zero] @[simp] lemma expect_dite_eq (i : ι) (f : ∀ j, i = j → M) : - 𝔼 j ∈ s, (if h : i = j then f j h else 0) = if i ∈ s then f i rfl /ℚ s.card else 0 := by + 𝔼 j ∈ s, (if h : i = j then f j h else 0) = if i ∈ s then f i rfl /ℚ #s else 0 := by split_ifs <;> simp [expect, *] @[simp] lemma expect_dite_eq' (i : ι) (f : ∀ j, j = i → M) : - 𝔼 j ∈ s, (if h : j = i then f j h else 0) = if i ∈ s then f i rfl /ℚ s.card else 0 := by + 𝔼 j ∈ s, (if h : j = i then f j h else 0) = if i ∈ s then f i rfl /ℚ #s else 0 := by split_ifs <;> simp [expect, *] @[simp] lemma expect_ite_eq (i : ι) (f : ι → M) : - 𝔼 j ∈ s, (if i = j then f j else 0) = if i ∈ s then f i /ℚ s.card else 0 := by + 𝔼 j ∈ s, (if i = j then f j else 0) = if i ∈ s then f i /ℚ #s else 0 := by split_ifs <;> simp [expect, *] @[simp] lemma expect_ite_eq' (i : ι) (f : ι → M) : - 𝔼 j ∈ s, (if j = i then f j else 0) = if i ∈ s then f i /ℚ s.card else 0 := by + 𝔼 j ∈ s, (if j = i then f j else 0) = if i ∈ s then f i /ℚ #s else 0 := by split_ifs <;> simp [expect, *] end DecidableEq @@ -279,7 +279,7 @@ lemma _root_.map_expect {F : Type*} [FunLike F M N] [LinearMapClass F ℚ≥0 M g (𝔼 i ∈ s, f i) = 𝔼 i ∈ s, g (f i) := by simp only [expect, map_smul, map_natCast, map_sum] @[simp] -lemma card_smul_expect (s : Finset ι) (f : ι → M) : s.card • 𝔼 i ∈ s, f i = ∑ i ∈ s, f i := by +lemma card_smul_expect (s : Finset ι) (f : ι → M) : #s • 𝔼 i ∈ s, f i = ∑ i ∈ s, f i := by obtain rfl | hs := s.eq_empty_or_nonempty · simp · rw [expect, ← Nat.cast_smul_eq_nsmul ℚ≥0, smul_inv_smul₀] @@ -315,7 +315,7 @@ section Semiring variable [Semiring M] [Module ℚ≥0 M] @[simp] lemma card_mul_expect (s : Finset ι) (f : ι → M) : - s.card * 𝔼 i ∈ s, f i = ∑ i ∈ s, f i := by rw [← nsmul_eq_mul, card_smul_expect] + #s * 𝔼 i ∈ s, f i = ∑ i ∈ s, f i := by rw [← nsmul_eq_mul, card_smul_expect] @[simp] lemma _root_.Fintype.card_mul_expect [Fintype ι] (f : ι → M) : Fintype.card ι * 𝔼 i, f i = ∑ i, f i := Finset.card_mul_expect _ _ @@ -357,7 +357,7 @@ lemma expect_boole_mul' [Fintype ι] [Nonempty ι] [DecidableEq ι] (f : ι → simp_rw [@eq_comm _ _ i, expect_boole_mul] lemma expect_eq_sum_div_card (s : Finset ι) (f : ι → M) : - 𝔼 i ∈ s, f i = (∑ i ∈ s, f i) / s.card := by + 𝔼 i ∈ s, f i = (∑ i ∈ s, f i) / #s := by rw [expect, NNRat.smul_def, div_eq_inv_mul, NNRat.cast_inv, NNRat.cast_natCast] lemma _root_.Fintype.expect_eq_sum_div_card [Fintype ι] (f : ι → M) : diff --git a/Mathlib/Algebra/BigOperators/Group/Finset.lean b/Mathlib/Algebra/BigOperators/Group/Finset.lean index 1dbe92862bdb6..4caade7753513 100644 --- a/Mathlib/Algebra/BigOperators/Group/Finset.lean +++ b/Mathlib/Algebra/BigOperators/Group/Finset.lean @@ -416,7 +416,7 @@ theorem prod_union [DecidableEq α] (h : Disjoint s₁ s₂) : @[to_additive] theorem prod_filter_mul_prod_filter_not (s : Finset α) (p : α → Prop) [DecidablePred p] [∀ x, Decidable (¬p x)] (f : α → β) : - (∏ x ∈ s.filter p, f x) * ∏ x ∈ s.filter fun x => ¬p x, f x = ∏ x ∈ s, f x := by + (∏ x ∈ s with p x, f x) * ∏ x ∈ s with ¬p x, f x = ∏ x ∈ s, f x := by have := Classical.decEq α rw [← prod_union (disjoint_filter_filter_neg s s p), filter_union_filter_neg_eq] @@ -463,11 +463,11 @@ lemma prod_powerset_cons (ha : a ∉ s) (f : Finset α → β) : rw [prod_powerset_insert ha, prod_attach _ fun t ↦ f (insert a t)] /-- A product over `powerset s` is equal to the double product over sets of subsets of `s` with -`card s = k`, for `k = 1, ..., card s`. -/ +`#s = k`, for `k = 1, ..., #s`. -/ @[to_additive "A sum over `powerset s` is equal to the double sum over sets of subsets of `s` with - `card s = k`, for `k = 1, ..., card s`"] +`#s = k`, for `k = 1, ..., #s`"] lemma prod_powerset (s : Finset α) (f : Finset α → β) : - ∏ t ∈ powerset s, f t = ∏ j ∈ range (card s + 1), ∏ t ∈ powersetCard j s, f t := by + ∏ t ∈ powerset s, f t = ∏ j ∈ range (#s + 1), ∏ t ∈ powersetCard j s, f t := by rw [powerset_card_disjiUnion, prod_disjiUnion] end CommMonoid @@ -676,27 +676,27 @@ variable [DecidableEq κ] @[to_additive] lemma prod_fiberwise_eq_prod_filter (s : Finset ι) (t : Finset κ) (g : ι → κ) (f : ι → α) : - ∏ j ∈ t, ∏ i ∈ s.filter fun i ↦ g i = j, f i = ∏ i ∈ s.filter fun i ↦ g i ∈ t, f i := by + ∏ j ∈ t, ∏ i ∈ s with g i = j, f i = ∏ i ∈ s with g i ∈ t, f i := by rw [← prod_disjiUnion, disjiUnion_filter_eq] @[to_additive] lemma prod_fiberwise_eq_prod_filter' (s : Finset ι) (t : Finset κ) (g : ι → κ) (f : κ → α) : - ∏ j ∈ t, ∏ _i ∈ s.filter fun i ↦ g i = j, f j = ∏ i ∈ s.filter fun i ↦ g i ∈ t, f (g i) := by + ∏ j ∈ t, ∏ i ∈ s with g i = j, f j = ∏ i ∈ s with g i ∈ t, f (g i) := by calc - _ = ∏ j ∈ t, ∏ i ∈ s.filter fun i ↦ g i = j, f (g i) := + _ = ∏ j ∈ t, ∏ i ∈ s with g i = j, f (g i) := prod_congr rfl fun j _ ↦ prod_congr rfl fun i hi ↦ by rw [(mem_filter.1 hi).2] _ = _ := prod_fiberwise_eq_prod_filter _ _ _ _ @[to_additive] lemma prod_fiberwise_of_maps_to {g : ι → κ} (h : ∀ i ∈ s, g i ∈ t) (f : ι → α) : - ∏ j ∈ t, ∏ i ∈ s.filter fun i ↦ g i = j, f i = ∏ i ∈ s, f i := by + ∏ j ∈ t, ∏ i ∈ s with g i = j, f i = ∏ i ∈ s, f i := by rw [← prod_disjiUnion, disjiUnion_filter_eq_of_maps_to h] @[to_additive] lemma prod_fiberwise_of_maps_to' {g : ι → κ} (h : ∀ i ∈ s, g i ∈ t) (f : κ → α) : - ∏ j ∈ t, ∏ _i ∈ s.filter fun i ↦ g i = j, f j = ∏ i ∈ s, f (g i) := by + ∏ j ∈ t, ∏ i ∈ s with g i = j, f j = ∏ i ∈ s, f (g i) := by calc - _ = ∏ y ∈ t, ∏ x ∈ s.filter fun x ↦ g x = y, f (g x) := + _ = ∏ j ∈ t, ∏ i ∈ s with g i = j, f (g i) := prod_congr rfl fun y _ ↦ prod_congr rfl fun x hx ↦ by rw [(mem_filter.1 hx).2] _ = _ := prod_fiberwise_of_maps_to h _ @@ -704,12 +704,12 @@ variable [Fintype κ] @[to_additive] lemma prod_fiberwise (s : Finset ι) (g : ι → κ) (f : ι → α) : - ∏ j, ∏ i ∈ s.filter fun i ↦ g i = j, f i = ∏ i ∈ s, f i := + ∏ j, ∏ i ∈ s with g i = j, f i = ∏ i ∈ s, f i := prod_fiberwise_of_maps_to (fun _ _ ↦ mem_univ _) _ @[to_additive] lemma prod_fiberwise' (s : Finset ι) (g : ι → κ) (f : κ → α) : - ∏ j, ∏ _i ∈ s.filter fun i ↦ g i = j, f j = ∏ i ∈ s, f (g i) := + ∏ j, ∏ i ∈ s with g i = j, f j = ∏ i ∈ s, f (g i) := prod_fiberwise_of_maps_to' (fun _ _ ↦ mem_univ _) _ end bij @@ -759,15 +759,15 @@ theorem prod_finset_product_right' (r : Finset (α × γ)) (s : Finset γ) (t : prod_finset_product_right r s t h @[to_additive] -theorem prod_image' [DecidableEq α] {s : Finset γ} {g : γ → α} (h : γ → β) - (eq : ∀ c ∈ s, f (g c) = ∏ x ∈ s.filter fun c' => g c' = g c, h x) : - ∏ x ∈ s.image g, f x = ∏ x ∈ s, h x := +theorem prod_image' [DecidableEq α] {s : Finset ι} {g : ι → α} (h : ι → β) + (eq : ∀ i ∈ s, f (g i) = ∏ j ∈ s with g j = g i, h j) : + ∏ a ∈ s.image g, f a = ∏ i ∈ s, h i := calc - ∏ x ∈ s.image g, f x = ∏ x ∈ s.image g, ∏ x ∈ s.filter fun c' => g c' = x, h x := - (prod_congr rfl) fun _x hx => - let ⟨c, hcs, hc⟩ := mem_image.1 hx - hc ▸ eq c hcs - _ = ∏ x ∈ s, h x := prod_fiberwise_of_maps_to (fun _x => mem_image_of_mem g) _ + ∏ a ∈ s.image g, f a = ∏ a ∈ s.image g, ∏ j ∈ s with g j = a, h j := + (prod_congr rfl) fun _a hx => + let ⟨i, his, hi⟩ := mem_image.1 hx + hi ▸ eq i his + _ = ∏ i ∈ s, h i := prod_fiberwise_of_maps_to (fun _ => mem_image_of_mem g) _ @[to_additive] theorem prod_mul_distrib : ∏ x ∈ s, f x * g x = (∏ x ∈ s, f x) * ∏ x ∈ s, g x := @@ -836,7 +836,7 @@ theorem prod_hom_rel [CommMonoid γ] {r : β → γ → Prop} {f : α → β} {g @[to_additive] theorem prod_filter_of_ne {p : α → Prop} [DecidablePred p] (hp : ∀ x ∈ s, f x ≠ 1 → p x) : - ∏ x ∈ s.filter p, f x = ∏ x ∈ s, f x := + ∏ x ∈ s with p x, f x = ∏ x ∈ s, f x := (prod_subset (filter_subset _ _)) fun x => by classical rw [not_imp_comm, mem_filter] @@ -846,14 +846,14 @@ theorem prod_filter_of_ne {p : α → Prop} [DecidablePred p] (hp : ∀ x ∈ s, -- instance first; `{∀ x, Decidable (f x ≠ 1)}` doesn't work with `rw ← prod_filter_ne_one` @[to_additive] theorem prod_filter_ne_one (s : Finset α) [∀ x, Decidable (f x ≠ 1)] : - ∏ x ∈ s.filter fun x => f x ≠ 1, f x = ∏ x ∈ s, f x := + ∏ x ∈ s with f x ≠ 1, f x = ∏ x ∈ s, f x := prod_filter_of_ne fun _ _ => id @[to_additive] theorem prod_filter (p : α → Prop) [DecidablePred p] (f : α → β) : - ∏ a ∈ s.filter p, f a = ∏ a ∈ s, if p a then f a else 1 := + ∏ a ∈ s with p a, f a = ∏ a ∈ s, if p a then f a else 1 := calc - ∏ a ∈ s.filter p, f a = ∏ a ∈ s.filter p, if p a then f a else 1 := + ∏ a ∈ s with p a, f a = ∏ a ∈ s with p a, if p a then f a else 1 := prod_congr rfl fun a h => by rw [if_pos]; simpa using (mem_filter.1 h).2 _ = ∏ a ∈ s, if p a then f a else 1 := by { refine prod_subset (filter_subset _ s) fun x hs h => ?_ @@ -929,11 +929,11 @@ theorem prod_eq_mul {s : Finset α} {f : α → β} (a b : α) (hn : a ≠ b) prod_const_one -- Porting note: simpNF linter complains that LHS doesn't simplify, but it does -/-- A product over `s.subtype p` equals one over `s.filter p`. -/ +/-- A product over `s.subtype p` equals one over `{x ∈ s | p x}`. -/ @[to_additive (attr := simp, nolint simpNF) - "A sum over `s.subtype p` equals one over `s.filter p`."] +"A sum over `s.subtype p` equals one over `{x ∈ s | p x}`."] theorem prod_subtype_eq_prod_filter (f : α → β) {p : α → Prop} [DecidablePred p] : - ∏ x ∈ s.subtype p, f x = ∏ x ∈ s.filter p, f x := by + ∏ x ∈ s.subtype p, f x = ∏ x ∈ s with p x, f x := by conv_lhs => erw [← prod_map (s.subtype p) (Function.Embedding.subtype _) f] exact prod_congr (subtype_map _) fun x _hx => rfl @@ -983,12 +983,12 @@ theorem prod_subtype {p : α → Prop} {F : Fintype (Subtype p)} (s : Finset α) @[to_additive] lemma prod_preimage' (f : ι → κ) [DecidablePred (· ∈ Set.range f)] (s : Finset κ) (hf) (g : κ → β) : - ∏ x ∈ s.preimage f hf, g (f x) = ∏ x ∈ s.filter (· ∈ Set.range f), g x := by + ∏ x ∈ s.preimage f hf, g (f x) = ∏ x ∈ s with x ∈ Set.range f, g x := by classical calc ∏ x ∈ preimage s f hf, g (f x) = ∏ x ∈ image f (preimage s f hf), g x := Eq.symm <| prod_image <| by simpa only [mem_preimage, Set.InjOn] using hf - _ = ∏ x ∈ s.filter fun x => x ∈ Set.range f, g x := by rw [image_preimage] + _ = ∏ x ∈ s with x ∈ Set.range f, g x := by rw [image_preimage] @[to_additive] lemma prod_preimage (f : ι → κ) (s : Finset κ) (hf) (g : κ → β) @@ -1026,18 +1026,18 @@ theorem prod_congr_set {α : Type*} [CommMonoid α] {β : Type*} [Fintype β] (s theorem prod_apply_dite {s : Finset α} {p : α → Prop} {hp : DecidablePred p} [DecidablePred fun x => ¬p x] (f : ∀ x : α, p x → γ) (g : ∀ x : α, ¬p x → γ) (h : γ → β) : (∏ x ∈ s, h (if hx : p x then f x hx else g x hx)) = - (∏ x ∈ (s.filter p).attach, h (f x.1 <| by simpa using (mem_filter.mp x.2).2)) * - ∏ x ∈ (s.filter fun x => ¬p x).attach, h (g x.1 <| by simpa using (mem_filter.mp x.2).2) := + (∏ x : {x ∈ s | p x}, h (f x.1 <| by simpa using (mem_filter.mp x.2).2)) * + ∏ x : {x ∈ s | ¬p x}, h (g x.1 <| by simpa using (mem_filter.mp x.2).2) := calc (∏ x ∈ s, h (if hx : p x then f x hx else g x hx)) = - (∏ x ∈ s.filter p, h (if hx : p x then f x hx else g x hx)) * - ∏ x ∈ s.filter (¬p ·), h (if hx : p x then f x hx else g x hx) := + (∏ x ∈ s with p x, h (if hx : p x then f x hx else g x hx)) * + ∏ x ∈ s with ¬p x, h (if hx : p x then f x hx else g x hx) := (prod_filter_mul_prod_filter_not s p _).symm - _ = (∏ x ∈ (s.filter p).attach, h (if hx : p x.1 then f x.1 hx else g x.1 hx)) * - ∏ x ∈ (s.filter (¬p ·)).attach, h (if hx : p x.1 then f x.1 hx else g x.1 hx) := + _ = (∏ x : {x ∈ s | p x}, h (if hx : p x.1 then f x.1 hx else g x.1 hx)) * + ∏ x : {x ∈ s | ¬p x}, h (if hx : p x.1 then f x.1 hx else g x.1 hx) := congr_arg₂ _ (prod_attach _ _).symm (prod_attach _ _).symm - _ = (∏ x ∈ (s.filter p).attach, h (f x.1 <| by simpa using (mem_filter.mp x.2).2)) * - ∏ x ∈ (s.filter (¬p ·)).attach, h (g x.1 <| by simpa using (mem_filter.mp x.2).2) := + _ = (∏ x : {x ∈ s | p x}, h (f x.1 <| by simpa using (mem_filter.mp x.2).2)) * + ∏ x : {x ∈ s | ¬p x}, h (g x.1 <| by simpa using (mem_filter.mp x.2).2) := congr_arg₂ _ (prod_congr rfl fun x _hx ↦ congr_arg h (dif_pos <| by simpa using (mem_filter.mp x.2).2)) (prod_congr rfl fun x _hx => congr_arg h (dif_neg <| by simpa using (mem_filter.mp x.2).2)) @@ -1046,21 +1046,20 @@ theorem prod_apply_dite {s : Finset α} {p : α → Prop} {hp : DecidablePred p} theorem prod_apply_ite {s : Finset α} {p : α → Prop} {_hp : DecidablePred p} (f g : α → γ) (h : γ → β) : (∏ x ∈ s, h (if p x then f x else g x)) = - (∏ x ∈ s.filter p, h (f x)) * ∏ x ∈ s.filter fun x => ¬p x, h (g x) := + (∏ x ∈ s with p x, h (f x)) * ∏ x ∈ s with ¬p x, h (g x) := (prod_apply_dite _ _ _).trans <| congr_arg₂ _ (prod_attach _ (h ∘ f)) (prod_attach _ (h ∘ g)) @[to_additive] theorem prod_dite {s : Finset α} {p : α → Prop} {hp : DecidablePred p} (f : ∀ x : α, p x → β) (g : ∀ x : α, ¬p x → β) : ∏ x ∈ s, (if hx : p x then f x hx else g x hx) = - (∏ x ∈ (s.filter p).attach, f x.1 (by simpa using (mem_filter.mp x.2).2)) * - ∏ x ∈ (s.filter fun x => ¬p x).attach, g x.1 (by simpa using (mem_filter.mp x.2).2) := by + (∏ x : {x ∈ s | p x}, f x.1 (by simpa using (mem_filter.mp x.2).2)) * + ∏ x : {x ∈ s | ¬p x}, g x.1 (by simpa using (mem_filter.mp x.2).2) := by simp [prod_apply_dite _ _ fun x => x] @[to_additive] theorem prod_ite {s : Finset α} {p : α → Prop} {hp : DecidablePred p} (f g : α → β) : - ∏ x ∈ s, (if p x then f x else g x) = - (∏ x ∈ s.filter p, f x) * ∏ x ∈ s.filter fun x => ¬p x, g x := by + ∏ x ∈ s, (if p x then f x else g x) = (∏ x ∈ s with p x, f x) * ∏ x ∈ s with ¬p x, g x := by simp [prod_apply_ite _ _ fun x => x] @[to_additive] @@ -1219,7 +1218,7 @@ lemma prod_mulIndicator_subset (f : ι → β) {s t : Finset ι} (h : s ⊆ t) : @[to_additive] lemma prod_mulIndicator_eq_prod_filter (s : Finset ι) (f : ι → κ → β) (t : ι → Set κ) (g : ι → κ) [DecidablePred fun i ↦ g i ∈ t i] : - ∏ i ∈ s, mulIndicator (t i) (f i) (g i) = ∏ i ∈ s.filter fun i ↦ g i ∈ t i, f i (g i) := by + ∏ i ∈ s, mulIndicator (t i) (f i) (g i) = ∏ i ∈ s with g i ∈ t i, f i (g i) := by refine (prod_filter_mul_prod_filter_not s (fun i ↦ g i ∈ t i) _).symm.trans <| Eq.trans (congr_arg₂ (· * ·) ?_ ?_) (mul_one _) · exact prod_congr rfl fun x hx ↦ mulIndicator_of_mem (mem_filter.1 hx).2 _ @@ -1264,8 +1263,8 @@ theorem prod_bij_ne_one {s : Finset α} {t : Finset γ} {f : α → β} {g : γ ∏ x ∈ s, f x = ∏ x ∈ t, g x := by classical calc - ∏ x ∈ s, f x = ∏ x ∈ s.filter fun x => f x ≠ 1, f x := by rw [prod_filter_ne_one] - _ = ∏ x ∈ t.filter fun x => g x ≠ 1, g x := + ∏ x ∈ s, f x = ∏ x ∈ s with f x ≠ 1, f x := by rw [prod_filter_ne_one] + _ = ∏ x ∈ t with g x ≠ 1, g x := prod_bij (fun a ha => i a (mem_filter.mp ha).1 <| by simpa using (mem_filter.mp ha).2) ?_ ?_ ?_ ?_ _ = ∏ x ∈ t, g x := prod_filter_ne_one _ @@ -1376,7 +1375,7 @@ theorem prod_list_count_of_subset [DecidableEq α] [CommMonoid α] (m : List α) rw [count_eq_zero_of_not_mem hx, pow_zero] theorem sum_filter_count_eq_countP [DecidableEq α] (p : α → Prop) [DecidablePred p] (l : List α) : - ∑ x ∈ l.toFinset.filter p, l.count x = l.countP p := by + ∑ x ∈ l.toFinset with p x, l.count x = l.countP p := by simp [Finset.sum, sum_map_count_dedup_filter_eq_countP p l] open Multiset @@ -1484,19 +1483,19 @@ theorem sum_tsub_distrib [AddCommMonoid α] [PartialOrder α] [ExistsAddOfLE α] ∑ x ∈ s, (f x - g x) = ∑ x ∈ s, f x - ∑ x ∈ s, g x := sum_map_tsub _ hfg @[to_additive (attr := simp)] -theorem prod_const (b : β) : ∏ _x ∈ s, b = b ^ s.card := - (congr_arg _ <| s.val.map_const b).trans <| Multiset.prod_replicate s.card b +theorem prod_const (b : β) : ∏ _x ∈ s, b = b ^ #s := + (congr_arg _ <| s.val.map_const b).trans <| Multiset.prod_replicate #s b @[to_additive sum_eq_card_nsmul] -theorem prod_eq_pow_card {b : β} (hf : ∀ a ∈ s, f a = b) : ∏ a ∈ s, f a = b ^ s.card := +theorem prod_eq_pow_card {b : β} (hf : ∀ a ∈ s, f a = b) : ∏ a ∈ s, f a = b ^ #s := (prod_congr rfl hf).trans <| prod_const _ @[to_additive card_nsmul_add_sum] -theorem pow_card_mul_prod {b : β} : b ^ s.card * ∏ a ∈ s, f a = ∏ a ∈ s, b * f a := +theorem pow_card_mul_prod {b : β} : b ^ #s * ∏ a ∈ s, f a = ∏ a ∈ s, b * f a := (Finset.prod_const b).symm ▸ prod_mul_distrib.symm @[to_additive sum_add_card_nsmul] -theorem prod_mul_pow_card {b : β} : (∏ a ∈ s, f a) * b ^ s.card = ∏ a ∈ s, f a * b := +theorem prod_mul_pow_card {b : β} : (∏ a ∈ s, f a) * b ^ #s = ∏ a ∈ s, f a * b := (Finset.prod_const b).symm ▸ prod_mul_distrib.symm @[to_additive] @@ -1515,7 +1514,7 @@ lemma prod_pow_eq_pow_sum (s : Finset ι) (f : ι → ℕ) (a : β) : @[to_additive "A sum over `Finset.powersetCard` which only depends on the size of the sets is constant."] lemma prod_powersetCard (n : ℕ) (s : Finset α) (f : ℕ → β) : - ∏ t ∈ powersetCard n s, f t.card = f n ^ s.card.choose n := by + ∏ t ∈ powersetCard n s, f #t = f n ^ (#s).choose n := by rw [prod_eq_pow_card, card_powersetCard]; rintro a ha; rw [(mem_powersetCard.1 ha).2] @[to_additive] @@ -1568,7 +1567,7 @@ lemma prod_ninvolution (g : α → α) (hg₁ : ∀ a, f a * f (g a) = 1) (hg₂ @[to_additive "The sum of the composition of functions `f` and `g`, is the sum over `b ∈ s.image g` of `f b` times of the cardinality of the fibre of `b`. See also `Finset.sum_image`."] theorem prod_comp [DecidableEq γ] (f : γ → β) (g : α → γ) : - ∏ a ∈ s, f (g a) = ∏ b ∈ s.image g, f b ^ (s.filter fun a => g a = b).card := by + ∏ a ∈ s, f (g a) = ∏ b ∈ s.image g, f b ^ #{a ∈ s | g a = b} := by simp_rw [← prod_const, prod_fiberwise_of_maps_to' fun _ ↦ mem_image_of_mem _] @[to_additive] @@ -1611,14 +1610,14 @@ theorem dvd_prod_of_mem (f : α → β) {a : α} {s : Finset α} (ha : a ∈ s) /-- A product can be partitioned into a product of products, each equivalent under a setoid. -/ @[to_additive "A sum can be partitioned into a sum of sums, each equivalent under a setoid."] theorem prod_partition (R : Setoid α) [DecidableRel R.r] : - ∏ x ∈ s, f x = ∏ xbar ∈ s.image (Quotient.mk _), ∏ y ∈ s.filter (⟦·⟧ = xbar), f y := by + ∏ x ∈ s, f x = ∏ xbar ∈ s.image (Quotient.mk _), ∏ y ∈ s with ⟦y⟧ = xbar, f y := by refine (Finset.prod_image' f fun x _hx => ?_).symm rfl /-- If we can partition a product into subsets that cancel out, then the whole product cancels. -/ @[to_additive "If we can partition a sum into subsets that cancel out, then the whole sum cancels."] theorem prod_cancels_of_partition_cancels (R : Setoid α) [DecidableRel R] - (h : ∀ x ∈ s, ∏ a ∈ s.filter fun y => R y x, f a = 1) : ∏ x ∈ s, f x = 1 := by + (h : ∀ x ∈ s, ∏ a ∈ s with R a x, f a = 1) : ∏ x ∈ s, f x = 1 := by rw [prod_partition R, ← Finset.prod_eq_one] intro xbar xbar_in_s obtain ⟨x, x_in_s, rfl⟩ := mem_image.mp xbar_in_s @@ -1645,12 +1644,12 @@ theorem prod_update_of_mem [DecidableEq α] {s : Finset α} {i : α} (h : i ∈ do the terms in that product. -/ @[to_additive eq_of_card_le_one_of_sum_eq "If a sum of a `Finset` of size at most 1 has a given value, so do the terms in that sum."] -theorem eq_of_card_le_one_of_prod_eq {s : Finset α} (hc : s.card ≤ 1) {f : α → β} {b : β} +theorem eq_of_card_le_one_of_prod_eq {s : Finset α} (hc : #s ≤ 1) {f : α → β} {b : β} (h : ∏ x ∈ s, f x = b) : ∀ x ∈ s, f x = b := by intro x hx - by_cases hc0 : s.card = 0 + by_cases hc0 : #s = 0 · exact False.elim (card_ne_zero_of_mem hx hc0) - · have h1 : s.card = 1 := le_antisymm hc (Nat.one_le_of_lt (Nat.pos_of_ne_zero hc0)) + · have h1 : #s = 1 := le_antisymm hc (Nat.one_le_of_lt (Nat.pos_of_ne_zero hc0)) rw [card_eq_one] at h1 obtain ⟨x2, hx2⟩ := h1 rw [hx2, mem_singleton] at hx @@ -1781,19 +1780,18 @@ lemma prod_sdiff_ne_prod_sdiff_iff : end CancelCommMonoid -theorem card_eq_sum_ones (s : Finset α) : s.card = ∑ _ ∈ s, 1 := by simp +theorem card_eq_sum_ones (s : Finset α) : #s = ∑ _ ∈ s, 1 := by simp -theorem sum_const_nat {m : ℕ} {f : α → ℕ} (h₁ : ∀ x ∈ s, f x = m) : - ∑ x ∈ s, f x = card s * m := by +theorem sum_const_nat {m : ℕ} {f : α → ℕ} (h₁ : ∀ x ∈ s, f x = m) : ∑ x ∈ s, f x = #s * m := by rw [← Nat.nsmul_eq_mul, ← sum_const] apply sum_congr rfl h₁ lemma sum_card_fiberwise_eq_card_filter {κ : Type*} [DecidableEq κ] (s : Finset ι) (t : Finset κ) - (g : ι → κ) : ∑ j ∈ t, (s.filter fun i ↦ g i = j).card = (s.filter fun i ↦ g i ∈ t).card := by + (g : ι → κ) : ∑ j ∈ t, #{i ∈ s | g i = j} = #{i ∈ s | g i ∈ t} := by simpa only [card_eq_sum_ones] using sum_fiberwise_eq_sum_filter _ _ _ _ -lemma card_filter (p) [DecidablePred p] (s : Finset α) : - (filter p s).card = ∑ a ∈ s, ite (p a) 1 0 := by simp [sum_ite] +lemma card_filter (p) [DecidablePred p] (s : Finset ι) : + #{i ∈ s | p i} = ∑ i ∈ s, ite (p i) 1 0 := by simp [sum_ite] section Opposite @@ -1853,37 +1851,33 @@ end CommGroup @[simp] theorem card_sigma {σ : α → Type*} (s : Finset α) (t : ∀ a, Finset (σ a)) : - card (s.sigma t) = ∑ a ∈ s, card (t a) := + #(s.sigma t) = ∑ a ∈ s, #(t a) := Multiset.card_sigma _ _ @[simp] theorem card_disjiUnion (s : Finset α) (t : α → Finset β) (h) : - (s.disjiUnion t h).card = s.sum fun i => (t i).card := + #(s.disjiUnion t h) = ∑ a ∈ s, #(t a) := Multiset.card_bind _ _ theorem card_biUnion [DecidableEq β] {s : Finset α} {t : α → Finset β} - (h : ∀ x ∈ s, ∀ y ∈ s, x ≠ y → Disjoint (t x) (t y)) : - (s.biUnion t).card = ∑ u ∈ s, card (t u) := - calc - (s.biUnion t).card = ∑ _ ∈ s.biUnion t, 1 := card_eq_sum_ones _ - _ = ∑ a ∈ s, ∑ _i ∈ t a, 1 := Finset.sum_biUnion h - _ = ∑ u ∈ s, card (t u) := by simp_rw [card_eq_sum_ones] + (h : ∀ x ∈ s, ∀ y ∈ s, x ≠ y → Disjoint (t x) (t y)) : #(s.biUnion t) = ∑ u ∈ s, #(t u) := by + simpa using sum_biUnion h (β := ℕ) (f := 1) theorem card_biUnion_le [DecidableEq β] {s : Finset α} {t : α → Finset β} : - (s.biUnion t).card ≤ ∑ a ∈ s, (t a).card := + #(s.biUnion t) ≤ ∑ a ∈ s, #(t a) := haveI := Classical.decEq α Finset.induction_on s (by simp) fun a s has ih => calc - ((insert a s).biUnion t).card ≤ (t a).card + (s.biUnion t).card := by - { rw [biUnion_insert]; exact Finset.card_union_le _ _ } - _ ≤ ∑ a ∈ insert a s, card (t a) := by rw [sum_insert has]; exact Nat.add_le_add_left ih _ + #((insert a s).biUnion t) ≤ #(t a) + #(s.biUnion t) := by + rw [biUnion_insert]; exact card_union_le .. + _ ≤ ∑ a ∈ insert a s, #(t a) := by rw [sum_insert has]; exact Nat.add_le_add_left ih _ theorem card_eq_sum_card_fiberwise [DecidableEq β] {f : α → β} {s : Finset α} {t : Finset β} - (H : ∀ x ∈ s, f x ∈ t) : s.card = ∑ a ∈ t, (s.filter fun x => f x = a).card := by + (H : ∀ x ∈ s, f x ∈ t) : #s = ∑ b ∈ t, #{a ∈ s | f a = b} := by simp only [card_eq_sum_ones, sum_fiberwise_of_maps_to H] theorem card_eq_sum_card_image [DecidableEq β] (f : α → β) (s : Finset α) : - s.card = ∑ a ∈ s.image f, (s.filter fun x => f x = a).card := + #s = ∑ b ∈ s.image f, #{a ∈ s | f a = b} := card_eq_sum_card_fiberwise fun _ => mem_image_of_mem _ theorem mem_sum {f : α → Multiset β} (s : Finset α) (b : β) : @@ -2265,8 +2259,7 @@ end AddCommMonoid theorem Finset.sum_sym2_filter_not_isDiag {ι α} [LinearOrder ι] [AddCommMonoid α] (s : Finset ι) (p : Sym2 ι → α) : - ∑ i in s.sym2.filter (¬ ·.IsDiag), p i = - ∑ i in s.offDiag.filter (fun i => i.1 < i.2), p s(i.1, i.2) := by + ∑ i ∈ s.sym2 with ¬ i.IsDiag, p i = ∑ i ∈ s.offDiag with i.1 < i.2, p s(i.1, i.2) := by rw [Finset.offDiag_filter_lt_eq_filter_le] conv_rhs => rw [← Finset.sum_subtype_eq_sum_filter] refine (Finset.sum_equiv Sym2.sortEquiv.symm ?_ ?_).symm diff --git a/Mathlib/Algebra/BigOperators/Ring.lean b/Mathlib/Algebra/BigOperators/Ring.lean index ffe59a97cc2ff..650b3b139aa99 100644 --- a/Mathlib/Algebra/BigOperators/Ring.lean +++ b/Mathlib/Algebra/BigOperators/Ring.lean @@ -26,11 +26,11 @@ section AddCommMonoidWithOne variable [AddCommMonoidWithOne α] lemma natCast_card_filter (p) [DecidablePred p] (s : Finset ι) : - ((filter p s).card : α) = ∑ a ∈ s, if p a then (1 : α) else 0 := by + (#{x ∈ s | p x} : α) = ∑ a ∈ s, if p a then (1 : α) else 0 := by rw [sum_ite, sum_const_zero, add_zero, sum_const, nsmul_one] @[simp] lemma sum_boole (p) [DecidablePred p] (s : Finset ι) : - (∑ x ∈ s, if p x then 1 else 0 : α) = (s.filter p).card := + (∑ x ∈ s, if p x then 1 else 0 : α) = #{x ∈ s | p x} := (natCast_card_filter _ _).symm end AddCommMonoidWithOne @@ -164,7 +164,7 @@ theorem prod_add (f g : ι → α) (s : Finset ι) : ∏ a ∈ s.attach, if p a.1 a.2 then f a.1 else g a.1 := prod_sum _ _ _ _ = ∑ t ∈ s.powerset, (∏ a ∈ t, f a) * ∏ a ∈ s \ t, g a := sum_bij' - (fun f _ ↦ s.filter fun a ↦ ∃ h : a ∈ s, f a h) + (fun f _ ↦ {a ∈ s | ∃ h : a ∈ s, f a h}) (fun t _ a _ => a ∈ t) (by simp) (by simp [Classical.em]) @@ -183,8 +183,7 @@ end DecidableEq theorem prod_add_ordered [LinearOrder ι] (s : Finset ι) (f g : ι → α) : ∏ i ∈ s, (f i + g i) = (∏ i ∈ s, f i) + - ∑ i ∈ s, - g i * (∏ j ∈ s.filter (· < i), (f j + g j)) * ∏ j ∈ s.filter fun j => i < j, f j := by + ∑ i ∈ s, g i * (∏ j ∈ s with j < i, (f j + g j)) * ∏ j ∈ s with i < j, f j := by refine Finset.induction_on_max s (by simp) ?_ clear s intro a s ha ihs @@ -202,21 +201,21 @@ theorem prod_add_ordered [LinearOrder ι] (s : Finset ι) (f g : ι → α) : mul_left_comm] exact mt (fun ha => (mem_filter.1 ha).1) ha' -/-- Summing `a^s.card * b^(n-s.card)` over all finite subsets `s` of a `Finset` -gives `(a + b)^s.card`. -/ +/-- Summing `a ^ #t * b ^ (n - #t)` over all finite subsets `t` of a finset `s` +gives `(a + b) ^ #s`. -/ theorem sum_pow_mul_eq_add_pow (a b : α) (s : Finset ι) : - (∑ t ∈ s.powerset, a ^ t.card * b ^ (s.card - t.card)) = (a + b) ^ s.card := by + (∑ t ∈ s.powerset, a ^ #t * b ^ (#s - #t)) = (a + b) ^ #s := by classical rw [← prod_const, prod_add] refine Finset.sum_congr rfl fun t ht => ?_ rw [prod_const, prod_const, ← card_sdiff (mem_powerset.1 ht)] -/-- Summing `a^s.card * b^(n-s.card)` over all finite subsets `s` of a fintype of cardinality `n` +/-- Summing `a^#s * b^(n-#s)` over all finite subsets `s` of a fintype of cardinality `n` gives `(a + b)^n`. The "good" proof involves expanding along all coordinates using the fact that `x^n` is multilinear, but multilinear maps are only available now over rings, so we give instead a proof reducing to the usual binomial theorem to have a result over semirings. -/ lemma _root_.Fintype.sum_pow_mul_eq_add_pow (ι : Type*) [Fintype ι] (a b : α) : - ∑ s : Finset ι, a ^ s.card * b ^ (Fintype.card ι - s.card) = (a + b) ^ Fintype.card ι := + ∑ s : Finset ι, a ^ #s * b ^ (Fintype.card ι - #s) = (a + b) ^ Fintype.card ι := Finset.sum_pow_mul_eq_add_pow _ _ _ @[norm_cast] @@ -232,8 +231,7 @@ variable [CommRing α] lemma prod_sub_ordered [LinearOrder ι] (s : Finset ι) (f g : ι → α) : ∏ i ∈ s, (f i - g i) = (∏ i ∈ s, f i) - - ∑ i ∈ s, - g i * (∏ j ∈ s.filter (· < i), (f j - g j)) * ∏ j ∈ s.filter fun j => i < j, f j := by + ∑ i ∈ s, g i * (∏ j ∈ s with j < i, (f j - g j)) * ∏ j ∈ s with i < j, f j := by simp only [sub_eq_add_neg] convert prod_add_ordered s f fun i => -g i simp @@ -241,7 +239,7 @@ lemma prod_sub_ordered [LinearOrder ι] (s : Finset ι) (f g : ι → α) : /-- `∏ i, (1 - f i) = 1 - ∑ i, f i * (∏ j < i, 1 - f j)`. This formula is useful in construction of a partition of unity from a collection of “bump” functions. -/ theorem prod_one_sub_ordered [LinearOrder ι] (s : Finset ι) (f : ι → α) : - ∏ i ∈ s, (1 - f i) = 1 - ∑ i ∈ s, f i * ∏ j ∈ s.filter (· < i), (1 - f j) := by + ∏ i ∈ s, (1 - f i) = 1 - ∑ i ∈ s, f i * ∏ j ∈ s with j < i, (1 - f j) := by rw [prod_sub_ordered] simp diff --git a/Mathlib/Algebra/GeomSum.lean b/Mathlib/Algebra/GeomSum.lean index 1e81b5455510c..0b34f0c5df0f1 100644 --- a/Mathlib/Algebra/GeomSum.lean +++ b/Mathlib/Algebra/GeomSum.lean @@ -134,7 +134,7 @@ theorem geom_sum₂_self {α : Type*} [CommRing α] (x : α) (n : ℕ) : _ = ∑ _i ∈ Finset.range n, x ^ (n - 1) := Finset.sum_congr rfl fun _ hi => congr_arg _ <| add_tsub_cancel_of_le <| Nat.le_sub_one_of_lt <| Finset.mem_range.1 hi - _ = (Finset.range n).card • x ^ (n - 1) := Finset.sum_const _ + _ = #(range n) • x ^ (n - 1) := sum_const _ _ = n * x ^ (n - 1) := by rw [Finset.card_range, nsmul_eq_mul] /-- $x^n-y^n = (x-y) \sum x^ky^{n-1-k}$ reformulated without `-` signs. -/ diff --git a/Mathlib/Algebra/Group/UniqueProds/Basic.lean b/Mathlib/Algebra/Group/UniqueProds/Basic.lean index bc6e50b4765f8..f43ae9290e9cb 100644 --- a/Mathlib/Algebra/Group/UniqueProds/Basic.lean +++ b/Mathlib/Algebra/Group/UniqueProds/Basic.lean @@ -52,6 +52,8 @@ assert_not_exists Algebra assert_not_exists Submodule assert_not_exists StarModule +open Finset + /-- Let `G` be a Type with multiplication, let `A B : Finset G` be finite subsets and let `a0 b0 : G` be two elements. `UniqueMul A B a0 b0` asserts `a0 * b0` can be written in at most one way as a product of an element of `A` and an element of `B`. -/ @@ -71,7 +73,7 @@ theorem of_subsingleton [Subsingleton G] : UniqueMul A B a0 b0 := by simp [UniqueMul, eq_iff_true_of_subsingleton] @[to_additive of_card_le_one] -theorem of_card_le_one (hA : A.Nonempty) (hB : B.Nonempty) (hA1 : A.card ≤ 1) (hB1 : B.card ≤ 1) : +theorem of_card_le_one (hA : A.Nonempty) (hB : B.Nonempty) (hA1 : #A ≤ 1) (hB1 : #B ≤ 1) : ∃ a ∈ A, ∃ b ∈ B, UniqueMul A B a b := by rw [Finset.card_le_one_iff] at hA1 hB1 obtain ⟨a, ha⟩ := hA; obtain ⟨b, hb⟩ := hB @@ -118,7 +120,7 @@ theorem iff_existsUnique (aA : a0 ∈ A) (bB : b0 ∈ B) : open Finset in @[to_additive iff_card_le_one] theorem iff_card_le_one [DecidableEq G] (ha0 : a0 ∈ A) (hb0 : b0 ∈ B) : - UniqueMul A B a0 b0 ↔ ((A ×ˢ B).filter (fun p ↦ p.1 * p.2 = a0 * b0)).card ≤ 1 := by + UniqueMul A B a0 b0 ↔ #{p ∈ A ×ˢ B | p.1 * p.2 = a0 * b0} ≤ 1 := by simp_rw [card_le_one_iff, mem_filter, mem_product] refine ⟨fun h p1 p2 ⟨⟨ha1, hb1⟩, he1⟩ ⟨⟨ha2, hb2⟩, he2⟩ ↦ ?_, fun h a b ha hb he ↦ ?_⟩ · have h1 := h ha1 hb1 he1; have h2 := h ha2 hb2 he2 @@ -210,7 +212,7 @@ open Finset in theorem of_image_filter [DecidableEq H] (f : G →ₙ* H) {A B : Finset G} {aG bG : G} {aH bH : H} (hae : f aG = aH) (hbe : f bG = bH) (huH : UniqueMul (A.image f) (B.image f) aH bH) - (huG : UniqueMul (A.filter (f · = aH)) (B.filter (f · = bH)) aG bG) : + (huG : UniqueMul {a ∈ A | f a = aH} {b ∈ B | f b = bH} aG bG) : UniqueMul A B aG bG := fun a b ha hb he ↦ by specialize huH (mem_image_of_mem _ ha) (mem_image_of_mem _ hb) rw [← map_mul, he, map_mul, hae, hbe] at huH @@ -245,7 +247,7 @@ of elements satisfying the `UniqueAdd` property. -/ class TwoUniqueSums (G) [Add G] : Prop where /-- For `A B` two finite sets whose product has cardinality at least 2, we can find at least two unique pairs. -/ - uniqueAdd_of_one_lt_card : ∀ {A B : Finset G}, 1 < A.card * B.card → + uniqueAdd_of_one_lt_card : ∀ {A B : Finset G}, 1 < #A * #B → ∃ p1 ∈ A ×ˢ B, ∃ p2 ∈ A ×ˢ B, p1 ≠ p2 ∧ UniqueAdd A B p1.1 p1.2 ∧ UniqueAdd A B p2.1 p2.2 /-- Let `G` be a Type with multiplication. `TwoUniqueProds G` asserts that any two non-empty @@ -254,16 +256,16 @@ of elements satisfying the `UniqueMul` property. -/ class TwoUniqueProds (G) [Mul G] : Prop where /-- For `A B` two finite sets whose product has cardinality at least 2, we can find at least two unique pairs. -/ - uniqueMul_of_one_lt_card : ∀ {A B : Finset G}, 1 < A.card * B.card → + uniqueMul_of_one_lt_card : ∀ {A B : Finset G}, 1 < #A * #B → ∃ p1 ∈ A ×ˢ B, ∃ p2 ∈ A ×ˢ B, p1 ≠ p2 ∧ UniqueMul A B p1.1 p1.2 ∧ UniqueMul A B p2.1 p2.2 attribute [to_additive] TwoUniqueProds @[to_additive] -lemma uniqueMul_of_twoUniqueMul {G} [Mul G] {A B : Finset G} (h : 1 < A.card * B.card → +lemma uniqueMul_of_twoUniqueMul {G} [Mul G] {A B : Finset G} (h : 1 < #A * #B → ∃ p1 ∈ A ×ˢ B, ∃ p2 ∈ A ×ˢ B, p1 ≠ p2 ∧ UniqueMul A B p1.1 p1.2 ∧ UniqueMul A B p2.1 p2.2) (hA : A.Nonempty) (hB : B.Nonempty) : ∃ a ∈ A, ∃ b ∈ B, UniqueMul A B a b := by - by_cases hc : A.card ≤ 1 ∧ B.card ≤ 1 + by_cases hc : #A ≤ 1 ∧ #B ≤ 1 · exact UniqueMul.of_card_le_one hA hB hc.1 hc.2 simp_rw [not_and_or, not_le] at hc rw [← Finset.card_pos] at hA hB @@ -388,7 +390,7 @@ open MulOpposite in obtain ⟨a, ha, b, hb, hu⟩ := uniqueMul_of_nonempty hc.1 hc.2.1 let C := A.map ⟨_, mul_right_injective a⁻¹⟩ -- C = a⁻¹A let D := B.map ⟨_, mul_left_injective b⁻¹⟩ -- D = Bb⁻¹ - have hcard : 1 < C.card ∨ 1 < D.card := by simp_rw [C, D, card_map]; exact hc.2.2 + have hcard : 1 < #C ∨ 1 < #D := by simp_rw [C, D, card_map]; exact hc.2.2 have hC : 1 ∈ C := mem_map.mpr ⟨a, ha, inv_mul_cancel a⟩ have hD : 1 ∈ D := mem_map.mpr ⟨b, hb, mul_inv_cancel b⟩ suffices ∃ c ∈ C, ∃ d ∈ D, (c ≠ 1 ∨ d ≠ 1) ∧ UniqueMul C D c d by @@ -439,13 +441,13 @@ open UniqueMul in let _ := isWellFounded_ssubset (α := ∀ i, G i) -- why need this? apply IsWellFounded.induction (· ⊂ ·) A; intro A ihA B hA apply IsWellFounded.induction (· ⊂ ·) B; intro B ihB hB - by_cases hc : A.card ≤ 1 ∧ B.card ≤ 1 + by_cases hc : #A ≤ 1 ∧ #B ≤ 1 · exact of_card_le_one hA hB hc.1 hc.2 simp_rw [not_and_or, not_le] at hc obtain ⟨i, hc⟩ := exists_or.mpr (hc.imp exists_of_one_lt_card_pi exists_of_one_lt_card_pi) obtain ⟨ai, hA, bi, hB, hi⟩ := uniqueMul_of_nonempty (hA.image (· i)) (hB.image (· i)) rw [mem_image, ← filter_nonempty_iff] at hA hB - let A' := A.filter (· i = ai); let B' := B.filter (· i = bi) + let A' := {a ∈ A | a i = ai}; let B' := {b ∈ B | b i = bi} obtain ⟨a0, ha0, b0, hb0, hu⟩ : ∃ a0 ∈ A', ∃ b0 ∈ B', UniqueMul A' B' a0 b0 := by rcases hc with hc | hc; · exact ihA A' (hc.2 ai) hA hB by_cases hA' : A' = A @@ -484,7 +486,7 @@ open Finset [TwoUniqueProds G] : TwoUniqueProds H where uniqueMul_of_one_lt_card {A B} hc := by classical - obtain hc' | hc' := lt_or_le 1 ((A.image f).card * (B.image f).card) + obtain hc' | hc' := lt_or_le 1 (#(A.image f) * #(B.image f)) · obtain ⟨⟨a1, b1⟩, h1, ⟨a2, b2⟩, h2, hne, hu1, hu2⟩ := uniqueMul_of_one_lt_card hc' simp_rw [mem_product, mem_image] at h1 h2 ⊢ obtain ⟨⟨a1, ha1, rfl⟩, b1, hb1, rfl⟩ := h1 @@ -536,11 +538,11 @@ instance instForall {ι} (G : ι → Type*) [∀ i, Mul (G i)] [∀ i, TwoUnique contrapose! hne; rw [Prod.mk.inj_iff] at hne ⊢ rw [← ha1.2, ← hb1.2, ← ha2.2, ← hb2.2, hne.1, hne.2]; exact ⟨rfl, rfl⟩ all_goals rcases hc with hc | hc; · exact ihA _ (hc.2 _) - · by_cases hA : A.filter (· i = p2.1) = A + · by_cases hA : {a ∈ A | a i = p2.1} = A · rw [hA] exact ihB _ (hc.2 _) · exact ihA _ ((A.filter_subset _).ssubset_of_ne hA) - · by_cases hA : A.filter (· i = p1.1) = A + · by_cases hA : {a ∈ A | a i = p1.1} = A · rw [hA] exact ihB _ (hc.2 _) · exact ihA _ ((A.filter_subset _).ssubset_of_ne hA) diff --git a/Mathlib/Algebra/Module/BigOperators.lean b/Mathlib/Algebra/Module/BigOperators.lean index 06ae5152f2d7f..9be3a61dd0e46 100644 --- a/Mathlib/Algebra/Module/BigOperators.lean +++ b/Mathlib/Algebra/Module/BigOperators.lean @@ -41,16 +41,16 @@ lemma Fintype.sum_smul_sum [Fintype α] [Fintype β] (f : α → R) (g : β → end AddCommMonoid -theorem Finset.cast_card [CommSemiring R] (s : Finset α) : (s.card : R) = ∑ _ ∈ s, 1 := by - rw [Finset.sum_const, Nat.smul_one_eq_cast] - open Finset +theorem Finset.cast_card [CommSemiring R] (s : Finset α) : (#s : R) = ∑ _ ∈ s, 1 := by + rw [Finset.sum_const, Nat.smul_one_eq_cast] + namespace Fintype variable [DecidableEq ι] [Fintype ι] [AddCommMonoid α] lemma sum_piFinset_apply (f : κ → α) (s : Finset κ) (i : ι) : - ∑ g ∈ piFinset fun _ : ι ↦ s, f (g i) = s.card ^ (card ι - 1) • ∑ b ∈ s, f b := by + ∑ g ∈ piFinset fun _ : ι ↦ s, f (g i) = #s ^ (card ι - 1) • ∑ b ∈ s, f b := by classical rw [Finset.sum_comp] simp only [eval_image_piFinset_const, card_filter_piFinset_const s, ite_smul, zero_smul, smul_sum, diff --git a/Mathlib/Algebra/MonoidAlgebra/Defs.lean b/Mathlib/Algebra/MonoidAlgebra/Defs.lean index 767b5198c52d8..2d212d128bc5d 100644 --- a/Mathlib/Algebra/MonoidAlgebra/Defs.lean +++ b/Mathlib/Algebra/MonoidAlgebra/Defs.lean @@ -397,14 +397,10 @@ theorem mul_apply_antidiagonal [Mul G] (f g : MonoidAlgebra k G) (x : G) (s : Fi calc (f * g) x = ∑ a₁ ∈ f.support, ∑ a₂ ∈ g.support, F (a₁, a₂) := mul_apply f g x _ = ∑ p ∈ f.support ×ˢ g.support, F p := by rw [Finset.sum_product] - _ = ∑ p ∈ (f.support ×ˢ g.support).filter fun p : G × G => p.1 * p.2 = x, f p.1 * g p.2 := + _ = ∑ p ∈ f.support ×ˢ g.support with p.1 * p.2 = x, f p.1 * g p.2 := (Finset.sum_filter _ _).symm - _ = ∑ p ∈ s.filter fun p : G × G => p.1 ∈ f.support ∧ p.2 ∈ g.support, f p.1 * g p.2 := - (sum_congr - (by - ext - simp only [mem_filter, mem_product, hs, and_comm]) - fun _ _ => rfl) + _ = ∑ p ∈ s with p.1 ∈ f.support ∧ p.2 ∈ g.support, f p.1 * g p.2 := by + congr! 1; ext; simp only [mem_filter, mem_product, hs, and_comm] _ = ∑ p ∈ s, f p.1 * g p.2 := sum_subset (filter_subset _ _) fun p hps hp => by simp only [mem_filter, mem_support_iff, not_and, Classical.not_not] at hp ⊢ diff --git a/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean b/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean index f799eabddb523..c0119e6787e06 100644 --- a/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean +++ b/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean @@ -175,38 +175,36 @@ lemma mul_le_prod {i j : ι} (hf : ∀ i ∈ s, 1 ≤ f i) (hi : i ∈ s) (hj : @[to_additive sum_le_card_nsmul] theorem prod_le_pow_card (s : Finset ι) (f : ι → N) (n : N) (h : ∀ x ∈ s, f x ≤ n) : - s.prod f ≤ n ^ s.card := by + s.prod f ≤ n ^ #s := by refine (Multiset.prod_le_pow_card (s.val.map f) n ?_).trans ?_ · simpa using h · simp @[to_additive card_nsmul_le_sum] theorem pow_card_le_prod (s : Finset ι) (f : ι → N) (n : N) (h : ∀ x ∈ s, n ≤ f x) : - n ^ s.card ≤ s.prod f := @Finset.prod_le_pow_card _ Nᵒᵈ _ _ _ _ h + n ^ #s ≤ s.prod f := @Finset.prod_le_pow_card _ Nᵒᵈ _ _ _ _ h theorem card_biUnion_le_card_mul [DecidableEq β] (s : Finset ι) (f : ι → Finset β) (n : ℕ) - (h : ∀ a ∈ s, (f a).card ≤ n) : (s.biUnion f).card ≤ s.card * n := + (h : ∀ a ∈ s, #(f a) ≤ n) : #(s.biUnion f) ≤ #s * n := card_biUnion_le.trans <| sum_le_card_nsmul _ _ _ h variable {ι' : Type*} [DecidableEq ι'] --- Porting note: Mathport warning: expanding binder collection (y «expr ∉ » t) @[to_additive sum_fiberwise_le_sum_of_sum_fiber_nonneg] theorem prod_fiberwise_le_prod_of_one_le_prod_fiber' {t : Finset ι'} {g : ι → ι'} {f : ι → N} - (h : ∀ y ∉ t, (1 : N) ≤ ∏ x ∈ s.filter fun x ↦ g x = y, f x) : - (∏ y ∈ t, ∏ x ∈ s.filter fun x ↦ g x = y, f x) ≤ ∏ x ∈ s, f x := + (h : ∀ y ∉ t, (1 : N) ≤ ∏ x ∈ s with g x = y, f x) : + (∏ y ∈ t, ∏ x ∈ s with g x = y, f x) ≤ ∏ x ∈ s, f x := calc - (∏ y ∈ t, ∏ x ∈ s.filter fun x ↦ g x = y, f x) ≤ - ∏ y ∈ t ∪ s.image g, ∏ x ∈ s.filter fun x ↦ g x = y, f x := + (∏ y ∈ t, ∏ x ∈ s with g x = y, f x) ≤ + ∏ y ∈ t ∪ s.image g, ∏ x ∈ s with g x = y, f x := prod_le_prod_of_subset_of_one_le' subset_union_left fun y _ ↦ h y _ = ∏ x ∈ s, f x := prod_fiberwise_of_maps_to (fun _ hx ↦ mem_union.2 <| Or.inr <| mem_image_of_mem _ hx) _ --- Porting note: Mathport warning: expanding binder collection (y «expr ∉ » t) @[to_additive sum_le_sum_fiberwise_of_sum_fiber_nonpos] theorem prod_le_prod_fiberwise_of_prod_fiber_le_one' {t : Finset ι'} {g : ι → ι'} {f : ι → N} - (h : ∀ y ∉ t, ∏ x ∈ s.filter fun x ↦ g x = y, f x ≤ 1) : - ∏ x ∈ s, f x ≤ ∏ y ∈ t, ∏ x ∈ s.filter fun x ↦ g x = y, f x := + (h : ∀ y ∉ t, ∏ x ∈ s with g x = y, f x ≤ 1) : + ∏ x ∈ s, f x ≤ ∏ y ∈ t, ∏ x ∈ s with g x = y, f x := @prod_fiberwise_le_prod_of_one_le_prod_fiber' _ Nᵒᵈ _ _ _ _ _ _ _ h end OrderedCommMonoid @@ -247,27 +245,26 @@ section Pigeonhole variable [DecidableEq β] theorem card_le_mul_card_image_of_maps_to {f : α → β} {s : Finset α} {t : Finset β} - (Hf : ∀ a ∈ s, f a ∈ t) (n : ℕ) (hn : ∀ a ∈ t, (s.filter fun x ↦ f x = a).card ≤ n) : - s.card ≤ n * t.card := + (Hf : ∀ a ∈ s, f a ∈ t) (n : ℕ) (hn : ∀ b ∈ t, #{a ∈ s | f a = b} ≤ n) : #s ≤ n * #t := calc - s.card = ∑ a ∈ t, (s.filter fun x ↦ f x = a).card := card_eq_sum_card_fiberwise Hf - _ ≤ ∑ _a ∈ t, n := sum_le_sum hn + #s = ∑ b ∈ t, #{a ∈ s | f a = b} := card_eq_sum_card_fiberwise Hf + _ ≤ ∑ _b ∈ t, n := sum_le_sum hn _ = _ := by simp [mul_comm] theorem card_le_mul_card_image {f : α → β} (s : Finset α) (n : ℕ) - (hn : ∀ a ∈ s.image f, (s.filter fun x ↦ f x = a).card ≤ n) : s.card ≤ n * (s.image f).card := + (hn : ∀ b ∈ s.image f, #{a ∈ s | f a = b} ≤ n) : #s ≤ n * #(s.image f) := card_le_mul_card_image_of_maps_to (fun _ ↦ mem_image_of_mem _) n hn theorem mul_card_image_le_card_of_maps_to {f : α → β} {s : Finset α} {t : Finset β} - (Hf : ∀ a ∈ s, f a ∈ t) (n : ℕ) (hn : ∀ a ∈ t, n ≤ (s.filter fun x ↦ f x = a).card) : - n * t.card ≤ s.card := + (Hf : ∀ a ∈ s, f a ∈ t) (n : ℕ) (hn : ∀ b ∈ t, n ≤ #{a ∈ s | f a = b}) : + n * #t ≤ #s := calc - n * t.card = ∑ _a ∈ t, n := by simp [mul_comm] - _ ≤ ∑ a ∈ t, (s.filter fun x ↦ f x = a).card := sum_le_sum hn - _ = s.card := by rw [← card_eq_sum_card_fiberwise Hf] + n * #t = ∑ _a ∈ t, n := by simp [mul_comm] + _ ≤ ∑ b ∈ t, #{a ∈ s | f a = b} := sum_le_sum hn + _ = #s := by rw [← card_eq_sum_card_fiberwise Hf] theorem mul_card_image_le_card {f : α → β} (s : Finset α) (n : ℕ) - (hn : ∀ a ∈ s.image f, n ≤ (s.filter fun x ↦ f x = a).card) : n * (s.image f).card ≤ s.card := + (hn : ∀ b ∈ s.image f, n ≤ #{a ∈ s | f a = b}) : n * #(s.image f) ≤ #s := mul_card_image_le_card_of_maps_to (fun _ ↦ mem_image_of_mem _) n hn end Pigeonhole @@ -278,56 +275,52 @@ variable [DecidableEq α] {s : Finset α} {B : Finset (Finset α)} {n : ℕ} /-- If every element belongs to at most `n` Finsets, then the sum of their sizes is at most `n` times how many they are. -/ -theorem sum_card_inter_le (h : ∀ a ∈ s, (B.filter (a ∈ ·)).card ≤ n) : - (∑ t ∈ B, (s ∩ t).card) ≤ s.card * n := by +theorem sum_card_inter_le (h : ∀ a ∈ s, #{b ∈ B | a ∈ b} ≤ n) : (∑ t ∈ B, #(s ∩ t)) ≤ #s * n := by refine le_trans ?_ (s.sum_le_card_nsmul _ _ h) simp_rw [← filter_mem_eq_inter, card_eq_sum_ones, sum_filter] exact sum_comm.le /-- If every element belongs to at most `n` Finsets, then the sum of their sizes is at most `n` times how many they are. -/ -theorem sum_card_le [Fintype α] (h : ∀ a, (B.filter (a ∈ ·)).card ≤ n) : - ∑ s ∈ B, s.card ≤ Fintype.card α * n := +lemma sum_card_le [Fintype α] (h : ∀ a, #{b ∈ B | a ∈ b} ≤ n) : ∑ s ∈ B, #s ≤ Fintype.card α * n := calc - ∑ s ∈ B, s.card = ∑ s ∈ B, (univ ∩ s).card := by simp_rw [univ_inter] + ∑ s ∈ B, #s = ∑ s ∈ B, #(univ ∩ s) := by simp_rw [univ_inter] _ ≤ Fintype.card α * n := sum_card_inter_le fun a _ ↦ h a /-- If every element belongs to at least `n` Finsets, then the sum of their sizes is at least `n` times how many they are. -/ -theorem le_sum_card_inter (h : ∀ a ∈ s, n ≤ (B.filter (a ∈ ·)).card) : - s.card * n ≤ ∑ t ∈ B, (s ∩ t).card := by +theorem le_sum_card_inter (h : ∀ a ∈ s, n ≤ #{b ∈ B | a ∈ b}) : #s * n ≤ ∑ t ∈ B, #(s ∩ t) := by apply (s.card_nsmul_le_sum _ _ h).trans simp_rw [← filter_mem_eq_inter, card_eq_sum_ones, sum_filter] exact sum_comm.le /-- If every element belongs to at least `n` Finsets, then the sum of their sizes is at least `n` times how many they are. -/ -theorem le_sum_card [Fintype α] (h : ∀ a, n ≤ (B.filter (a ∈ ·)).card) : - Fintype.card α * n ≤ ∑ s ∈ B, s.card := +theorem le_sum_card [Fintype α] (h : ∀ a, n ≤ #{b ∈ B | a ∈ b}) : + Fintype.card α * n ≤ ∑ s ∈ B, #s := calc - Fintype.card α * n ≤ ∑ s ∈ B, (univ ∩ s).card := le_sum_card_inter fun a _ ↦ h a - _ = ∑ s ∈ B, s.card := by simp_rw [univ_inter] + Fintype.card α * n ≤ ∑ s ∈ B, #(univ ∩ s) := le_sum_card_inter fun a _ ↦ h a + _ = ∑ s ∈ B, #s := by simp_rw [univ_inter] /-- If every element belongs to exactly `n` Finsets, then the sum of their sizes is `n` times how many they are. -/ -theorem sum_card_inter (h : ∀ a ∈ s, (B.filter (a ∈ ·)).card = n) : - (∑ t ∈ B, (s ∩ t).card) = s.card * n := +theorem sum_card_inter (h : ∀ a ∈ s, #{b ∈ B | a ∈ b} = n) : + (∑ t ∈ B, #(s ∩ t)) = #s * n := (sum_card_inter_le fun a ha ↦ (h a ha).le).antisymm (le_sum_card_inter fun a ha ↦ (h a ha).ge) /-- If every element belongs to exactly `n` Finsets, then the sum of their sizes is `n` times how many they are. -/ -theorem sum_card [Fintype α] (h : ∀ a, (B.filter (a ∈ ·)).card = n) : - ∑ s ∈ B, s.card = Fintype.card α * n := by +theorem sum_card [Fintype α] (h : ∀ a, #{b ∈ B | a ∈ b} = n) : + ∑ s ∈ B, #s = Fintype.card α * n := by simp_rw [Fintype.card, ← sum_card_inter fun a _ ↦ h a, univ_inter] theorem card_le_card_biUnion {s : Finset ι} {f : ι → Finset α} (hs : (s : Set ι).PairwiseDisjoint f) - (hf : ∀ i ∈ s, (f i).Nonempty) : s.card ≤ (s.biUnion f).card := by + (hf : ∀ i ∈ s, (f i).Nonempty) : #s ≤ #(s.biUnion f) := by rw [card_biUnion hs, card_eq_sum_ones] exact sum_le_sum fun i hi ↦ (hf i hi).card_pos theorem card_le_card_biUnion_add_card_fiber {s : Finset ι} {f : ι → Finset α} - (hs : (s : Set ι).PairwiseDisjoint f) : - s.card ≤ (s.biUnion f).card + (s.filter fun i ↦ f i = ∅).card := by + (hs : (s : Set ι).PairwiseDisjoint f) : #s ≤ #(s.biUnion f) + #{i ∈ s | f i = ∅} := by rw [← Finset.filter_card_add_filter_neg_card_eq_card fun i ↦ f i = ∅, add_comm] exact add_le_add_right @@ -337,7 +330,7 @@ theorem card_le_card_biUnion_add_card_fiber {s : Finset ι} {f : ι → Finset _ theorem card_le_card_biUnion_add_one {s : Finset ι} {f : ι → Finset α} (hf : Injective f) - (hs : (s : Set ι).PairwiseDisjoint f) : s.card ≤ (s.biUnion f).card + 1 := + (hs : (s : Set ι).PairwiseDisjoint f) : #s ≤ #(s.biUnion f) + 1 := (card_le_card_biUnion_add_card_fiber hs).trans <| add_le_add_left (card_le_one.2 fun _ hi _ hj ↦ hf <| (mem_filter.1 hi).2.trans (mem_filter.1 hj).2.symm) _ @@ -374,8 +367,7 @@ theorem prod_mono_set' (f : ι → M) : Monotone fun s ↦ ∏ x ∈ s, f x := f theorem prod_le_prod_of_ne_one' (h : ∀ x ∈ s, f x ≠ 1 → x ∈ t) : ∏ x ∈ s, f x ≤ ∏ x ∈ t, f x := by classical calc - ∏ x ∈ s, f x = (∏ x ∈ s.filter fun x ↦ f x = 1, f x) * - ∏ x ∈ s.filter fun x ↦ f x ≠ 1, f x := by + ∏ x ∈ s, f x = (∏ x ∈ s with f x = 1, f x) * ∏ x ∈ s with f x ≠ 1, f x := by rw [← prod_union, filter_union_filter_neg_eq] exact disjoint_filter.2 fun _ _ h n_h ↦ n_h h _ ≤ ∏ x ∈ t, f x := diff --git a/Mathlib/Algebra/Order/CauSeq/BigOperators.lean b/Mathlib/Algebra/Order/CauSeq/BigOperators.lean index 99ac7eb6f4dd3..ed1fb71669c94 100644 --- a/Mathlib/Algebra/Order/CauSeq/BigOperators.lean +++ b/Mathlib/Algebra/Order/CauSeq/BigOperators.lean @@ -121,9 +121,9 @@ theorem _root_.cauchy_product (ha : IsCauSeq abs fun m ↦ ∑ n ∈ range m, ab (by rw [← sum_mul, mul_comm]; gcongr) rw [sum_range_sub_sum_range (le_of_lt hNMK)] calc - (∑ i ∈ (range K).filter fun k ↦ max N M + 1 ≤ k, + (∑ i ∈ range K with max N M + 1 ≤ i, abv (f i) * abv ((∑ k ∈ range (K - i), g k) - ∑ k ∈ range K, g k)) ≤ - ∑ i ∈ (range K).filter fun k ↦ max N M + 1 ≤ k, abv (f i) * (2 * Q) := by + ∑ i ∈ range K with max N M + 1 ≤ i, abv (f i) * (2 * Q) := by gcongr rw [sub_eq_add_neg] refine le_trans (abv_add _ _ _) ?_ diff --git a/Mathlib/Algebra/Order/Chebyshev.lean b/Mathlib/Algebra/Order/Chebyshev.lean index 7987886b72b00..25c25f0275d3c 100644 --- a/Mathlib/Algebra/Order/Chebyshev.lean +++ b/Mathlib/Algebra/Order/Chebyshev.lean @@ -15,7 +15,7 @@ import Mathlib.Tactic.Positivity.Finset This file proves the Chebyshev sum inequality. -Chebyshev's inequality states `(∑ i ∈ s, f i) * (∑ i ∈ s, g i) ≤ s.card * ∑ i ∈ s, f i * g i` +Chebyshev's inequality states `(∑ i ∈ s, f i) * (∑ i ∈ s, g i) ≤ #s * ∑ i ∈ s, f i * g i` when `f g : ι → α` monovary, and the reverse inequality when `f` and `g` antivary. @@ -53,10 +53,10 @@ variable [LinearOrderedSemiring α] [ExistsAddOfLE α] [LinearOrderedCancelAddCo monotone/antitone), the scalar product of their sum is less than the size of the set times their scalar product. -/ theorem MonovaryOn.sum_smul_sum_le_card_smul_sum (hfg : MonovaryOn f g s) : - (∑ i ∈ s, f i) • ∑ i ∈ s, g i ≤ s.card • ∑ i ∈ s, f i • g i := by + (∑ i ∈ s, f i) • ∑ i ∈ s, g i ≤ #s • ∑ i ∈ s, f i • g i := by classical obtain ⟨σ, hσ, hs⟩ := s.countable_toSet.exists_cycleOn - rw [← card_range s.card, sum_smul_sum_eq_sum_perm hσ] + rw [← card_range #s, sum_smul_sum_eq_sum_perm hσ] exact sum_le_card_nsmul _ _ _ fun n _ ↦ hfg.sum_smul_comp_perm_le_sum_smul fun x hx ↦ hs fun h ↦ hx <| IsFixedPt.perm_pow h _ @@ -64,7 +64,7 @@ theorem MonovaryOn.sum_smul_sum_le_card_smul_sum (hfg : MonovaryOn f g s) : other is antitone), the scalar product of their sum is less than the size of the set times their scalar product. -/ theorem AntivaryOn.card_smul_sum_le_sum_smul_sum (hfg : AntivaryOn f g s) : - s.card • ∑ i ∈ s, f i • g i ≤ (∑ i ∈ s, f i) • ∑ i ∈ s, g i := + #s • ∑ i ∈ s, f i • g i ≤ (∑ i ∈ s, f i) • ∑ i ∈ s, g i := hfg.dual_right.sum_smul_sum_le_card_smul_sum variable [Fintype ι] @@ -99,7 +99,7 @@ variable [LinearOrderedSemiring α] [ExistsAddOfLE α] {s : Finset ι} {σ : Per monotone/antitone), the product of their sum is less than the size of the set times their scalar product. -/ theorem MonovaryOn.sum_mul_sum_le_card_mul_sum (hfg : MonovaryOn f g s) : - (∑ i ∈ s, f i) * ∑ i ∈ s, g i ≤ s.card * ∑ i ∈ s, f i * g i := by + (∑ i ∈ s, f i) * ∑ i ∈ s, g i ≤ #s * ∑ i ∈ s, f i * g i := by rw [← nsmul_eq_mul] exact hfg.sum_smul_sum_le_card_smul_sum @@ -107,29 +107,29 @@ theorem MonovaryOn.sum_mul_sum_le_card_mul_sum (hfg : MonovaryOn f g s) : other is antitone), the product of their sum is greater than the size of the set times their scalar product. -/ theorem AntivaryOn.card_mul_sum_le_sum_mul_sum (hfg : AntivaryOn f g s) : - (s.card : α) * ∑ i ∈ s, f i * g i ≤ (∑ i ∈ s, f i) * ∑ i ∈ s, g i := by + (#s : α) * ∑ i ∈ s, f i * g i ≤ (∑ i ∈ s, f i) * ∑ i ∈ s, g i := by rw [← nsmul_eq_mul] exact hfg.card_smul_sum_le_sum_smul_sum /-- Special case of **Jensen's inequality** for sums of powers. -/ lemma pow_sum_le_card_mul_sum_pow (hf : ∀ i ∈ s, 0 ≤ f i) : - ∀ n, (∑ i ∈ s, f i) ^ (n + 1) ≤ (s.card : α) ^ n * ∑ i ∈ s, f i ^ (n + 1) + ∀ n, (∑ i ∈ s, f i) ^ (n + 1) ≤ (#s : α) ^ n * ∑ i ∈ s, f i ^ (n + 1) | 0 => by simp | n + 1 => calc _ = (∑ i ∈ s, f i) ^ (n + 1) * ∑ i ∈ s, f i := by rw [pow_succ] - _ ≤ (s.card ^ n * ∑ i ∈ s, f i ^ (n + 1)) * ∑ i ∈ s, f i := by + _ ≤ (#s ^ n * ∑ i ∈ s, f i ^ (n + 1)) * ∑ i ∈ s, f i := by gcongr exacts [sum_nonneg hf, pow_sum_le_card_mul_sum_pow hf _] - _ = s.card ^ n * ((∑ i ∈ s, f i ^ (n + 1)) * ∑ i ∈ s, f i) := by rw [mul_assoc] - _ ≤ s.card ^ n * (s.card * ∑ i ∈ s, f i ^ (n + 1) * f i) := by + _ = #s ^ n * ((∑ i ∈ s, f i ^ (n + 1)) * ∑ i ∈ s, f i) := by rw [mul_assoc] + _ ≤ #s ^ n * (#s * ∑ i ∈ s, f i ^ (n + 1) * f i) := by gcongr _ * ?_ exact ((monovaryOn_self ..).pow_left₀ hf _).sum_mul_sum_le_card_mul_sum _ = _ := by simp_rw [← mul_assoc, ← pow_succ] /-- Special case of **Chebyshev's Sum Inequality** or the **Cauchy-Schwarz Inequality**: The square of the sum is less than the size of the set times the sum of the squares. -/ -theorem sq_sum_le_card_mul_sum_sq : (∑ i ∈ s, f i) ^ 2 ≤ s.card * ∑ i ∈ s, f i ^ 2 := by +theorem sq_sum_le_card_mul_sum_sq : (∑ i ∈ s, f i) ^ 2 ≤ #s * ∑ i ∈ s, f i ^ 2 := by simp_rw [sq] exact (monovaryOn_self _ _).sum_mul_sum_le_card_mul_sum @@ -155,16 +155,16 @@ variable [LinearOrderedSemifield α] [ExistsAddOfLE α] {s : Finset ι} {f : ι /-- Special case of **Jensen's inequality** for sums of powers. -/ lemma pow_sum_div_card_le_sum_pow (hf : ∀ i ∈ s, 0 ≤ f i) (n : ℕ) : - (∑ i ∈ s, f i) ^ (n + 1) / s.card ^ n ≤ ∑ i ∈ s, f i ^ (n + 1) := by + (∑ i ∈ s, f i) ^ (n + 1) / #s ^ n ≤ ∑ i ∈ s, f i ^ (n + 1) := by obtain rfl | hs := s.eq_empty_or_nonempty · simp rw [div_le_iff₀' (by positivity)] exact pow_sum_le_card_mul_sum_pow hf _ theorem sum_div_card_sq_le_sum_sq_div_card : - ((∑ i ∈ s, f i) / s.card) ^ 2 ≤ (∑ i ∈ s, f i ^ 2) / s.card := by + ((∑ i ∈ s, f i) / #s) ^ 2 ≤ (∑ i ∈ s, f i ^ 2) / #s := by obtain rfl | hs := s.eq_empty_or_nonempty · simp - rw [div_pow, div_le_div_iff (by positivity) (by positivity), sq (s.card : α), mul_left_comm, + rw [div_pow, div_le_div_iff (by positivity) (by positivity), sq (#s : α), mul_left_comm, ← mul_assoc] exact mul_le_mul_of_nonneg_right sq_sum_le_card_mul_sum_sq (by positivity) diff --git a/Mathlib/Algebra/Polynomial/Basic.lean b/Mathlib/Algebra/Polynomial/Basic.lean index b1d85a96da4bd..b3f8ed139f76f 100644 --- a/Mathlib/Algebra/Polynomial/Basic.lean +++ b/Mathlib/Algebra/Polynomial/Basic.lean @@ -60,7 +60,7 @@ structure Polynomial (R : Type*) [Semiring R] where ofFinsupp :: @[inherit_doc] scoped[Polynomial] notation:9000 R "[X]" => Polynomial R -open AddMonoidAlgebra +open AddMonoidAlgebra Finset open Finsupp hiding single open Function hiding Commute @@ -361,7 +361,7 @@ theorem support_eq_empty : p.support = ∅ ↔ p = 0 := by @[simp] lemma support_nonempty : p.support.Nonempty ↔ p ≠ 0 := Finset.nonempty_iff_ne_empty.trans support_eq_empty.not -theorem card_support_eq_zero : p.support.card = 0 ↔ p = 0 := by simp +theorem card_support_eq_zero : #p.support = 0 ↔ p = 0 := by simp /-- `monomial s a` is the monomial `a * X^s` -/ def monomial (n : ℕ) : R →ₗ[R] R[X] where diff --git a/Mathlib/Algebra/Polynomial/BigOperators.lean b/Mathlib/Algebra/Polynomial/BigOperators.lean index 909aed4291d6a..53e09909c6d90 100644 --- a/Mathlib/Algebra/Polynomial/BigOperators.lean +++ b/Mathlib/Algebra/Polynomial/BigOperators.lean @@ -214,7 +214,7 @@ theorem coeff_multiset_prod_of_natDegree_le (n : ℕ) (hl : ∀ p ∈ t, natDegr simpa using coeff_list_prod_of_natDegree_le _ _ hl theorem coeff_prod_of_natDegree_le (f : ι → R[X]) (n : ℕ) (h : ∀ p ∈ s, natDegree (f p) ≤ n) : - coeff (∏ i ∈ s, f i) (s.card * n) = ∏ i ∈ s, coeff (f i) n := by + coeff (∏ i ∈ s, f i) (#s * n) = ∏ i ∈ s, coeff (f i) n := by cases' s with l hl convert coeff_multiset_prod_of_natDegree_le (l.map f) n ?_ · simp @@ -266,8 +266,8 @@ theorem multiset_prod_X_sub_C_coeff_card_pred (t : Multiset R) (ht : 0 < Multise exact fun h => one_ne_zero <| h 1 ⟨_, ⟨x, hx, rfl⟩, natDegree_X_sub_C _⟩ congr; rw [natDegree_multiset_prod_of_monic] <;> · simp [natDegree_X_sub_C, monic_X_sub_C] -theorem prod_X_sub_C_coeff_card_pred (s : Finset ι) (f : ι → R) (hs : 0 < s.card) : - (∏ i ∈ s, (X - C (f i))).coeff (s.card - 1) = -∑ i ∈ s, f i := by +theorem prod_X_sub_C_coeff_card_pred (s : Finset ι) (f : ι → R) (hs : 0 < #s) : + (∏ i ∈ s, (X - C (f i))).coeff (#s - 1) = -∑ i ∈ s, f i := by simpa using multiset_prod_X_sub_C_coeff_card_pred (s.1.map f) (by simpa using hs) end CommRing diff --git a/Mathlib/Algebra/Polynomial/Coeff.lean b/Mathlib/Algebra/Polynomial/Coeff.lean index 3d6efd887e11e..82ae37b47a959 100644 --- a/Mathlib/Algebra/Polynomial/Coeff.lean +++ b/Mathlib/Algebra/Polynomial/Coeff.lean @@ -55,12 +55,12 @@ theorem support_smul [SMulZeroClass S R] (r : S) (p : R[X]) : simp [hi] open scoped Pointwise in -theorem card_support_mul_le : (p * q).support.card ≤ p.support.card * q.support.card := by - calc (p * q).support.card - _ = (p.toFinsupp * q.toFinsupp).support.card := by rw [← support_toFinsupp, toFinsupp_mul] - _ ≤ (p.toFinsupp.support + q.toFinsupp.support).card := +theorem card_support_mul_le : #(p * q).support ≤ #p.support * #q.support := by + calc #(p * q).support + _ = #(p.toFinsupp * q.toFinsupp).support := by rw [← support_toFinsupp, toFinsupp_mul] + _ ≤ #(p.toFinsupp.support + q.toFinsupp.support) := Finset.card_le_card (AddMonoidAlgebra.support_mul p.toFinsupp q.toFinsupp) - _ ≤ p.support.card * q.support.card := Finset.card_image₂_le .. + _ ≤ #p.support * #q.support := Finset.card_image₂_le .. /-- `Polynomial.sum` as a linear map. -/ @[simps] @@ -212,11 +212,11 @@ theorem support_trinomial {k m n : ℕ} (hkm : k < m) (hmn : m < n) {x y z : R} zero_add, Ne, hx, hy, hz, not_false_eq_true, and_true] theorem card_support_binomial {k m : ℕ} (h : k ≠ m) {x y : R} (hx : x ≠ 0) (hy : y ≠ 0) : - card (support (C x * X ^ k + C y * X ^ m)) = 2 := by + #(support (C x * X ^ k + C y * X ^ m)) = 2 := by rw [support_binomial h hx hy, card_insert_of_not_mem (mt mem_singleton.mp h), card_singleton] theorem card_support_trinomial {k m n : ℕ} (hkm : k < m) (hmn : m < n) {x y z : R} (hx : x ≠ 0) - (hy : y ≠ 0) (hz : z ≠ 0) : card (support (C x * X ^ k + C y * X ^ m + C z * X ^ n)) = 3 := by + (hy : y ≠ 0) (hz : z ≠ 0) : #(support (C x * X ^ k + C y * X ^ m + C z * X ^ n)) = 3 := by rw [support_trinomial hkm hmn hx hy hz, card_insert_of_not_mem (mt mem_insert.mp (not_or_intro hkm.ne (mt mem_singleton.mp (hkm.trans hmn).ne))), diff --git a/Mathlib/Algebra/Polynomial/Degree/Definitions.lean b/Mathlib/Algebra/Polynomial/Degree/Definitions.lean index 3a3fe3f22832f..80fccb38b65f8 100644 --- a/Mathlib/Algebra/Polynomial/Degree/Definitions.lean +++ b/Mathlib/Algebra/Polynomial/Degree/Definitions.lean @@ -414,11 +414,11 @@ theorem natDegree_X_le : (X : R[X]).natDegree ≤ 1 := theorem mem_support_C_mul_X_pow {n a : ℕ} {c : R} (h : a ∈ support (C c * X ^ n)) : a = n := mem_singleton.1 <| support_C_mul_X_pow' n c h -theorem card_support_C_mul_X_pow_le_one {c : R} {n : ℕ} : card (support (C c * X ^ n)) ≤ 1 := by +theorem card_support_C_mul_X_pow_le_one {c : R} {n : ℕ} : #(support (C c * X ^ n)) ≤ 1 := by rw [← card_singleton n] apply card_le_card (support_C_mul_X_pow' n c) -theorem card_supp_le_succ_natDegree (p : R[X]) : p.support.card ≤ p.natDegree + 1 := by +theorem card_supp_le_succ_natDegree (p : R[X]) : #p.support ≤ p.natDegree + 1 := by rw [← Finset.card_range (p.natDegree + 1)] exact Finset.card_le_card supp_subset_range_natDegree_succ @@ -888,13 +888,13 @@ theorem leadingCoeff_mul' (h : leadingCoeff p * leadingCoeff q ≠ 0) : rw [natDegree_mul' h, coeff_mul_degree_add_degree] rfl -theorem monomial_natDegree_leadingCoeff_eq_self (h : p.support.card ≤ 1) : +theorem monomial_natDegree_leadingCoeff_eq_self (h : #p.support ≤ 1) : monomial p.natDegree p.leadingCoeff = p := by classical rcases card_support_le_one_iff_monomial.1 h with ⟨n, a, rfl⟩ by_cases ha : a = 0 <;> simp [ha] -theorem C_mul_X_pow_eq_self (h : p.support.card ≤ 1) : C p.leadingCoeff * X ^ p.natDegree = p := by +theorem C_mul_X_pow_eq_self (h : #p.support ≤ 1) : C p.leadingCoeff * X ^ p.natDegree = p := by rw [C_mul_X_pow_eq_monomial, monomial_natDegree_leadingCoeff_eq_self h] theorem leadingCoeff_pow' : leadingCoeff p ^ n ≠ 0 → leadingCoeff (p ^ n) = leadingCoeff p ^ n := diff --git a/Mathlib/Algebra/Polynomial/EraseLead.lean b/Mathlib/Algebra/Polynomial/EraseLead.lean index 9f38cdcb5190b..3a13d453a7189 100644 --- a/Mathlib/Algebra/Polynomial/EraseLead.lean +++ b/Mathlib/Algebra/Polynomial/EraseLead.lean @@ -72,7 +72,7 @@ theorem self_sub_C_mul_X_pow {R : Type*} [Ring R] (f : R[X]) : f - C f.leadingCoeff * X ^ f.natDegree = f.eraseLead := by rw [C_mul_X_pow_eq_monomial, self_sub_monomial_natDegree_leadingCoeff] -theorem eraseLead_ne_zero (f0 : 2 ≤ f.support.card) : eraseLead f ≠ 0 := by +theorem eraseLead_ne_zero (f0 : 2 ≤ #f.support) : eraseLead f ≠ 0 := by rw [Ne, ← card_support_eq_zero, eraseLead_support] exact (zero_lt_one.trans_le <| (tsub_le_tsub_right f0 1).trans Finset.pred_card_le_card_erase).ne.symm @@ -89,13 +89,12 @@ theorem ne_natDegree_of_mem_eraseLead_support {a : ℕ} (h : a ∈ (eraseLead f) theorem natDegree_not_mem_eraseLead_support : f.natDegree ∉ (eraseLead f).support := fun h => ne_natDegree_of_mem_eraseLead_support h rfl -theorem eraseLead_support_card_lt (h : f ≠ 0) : (eraseLead f).support.card < f.support.card := by +theorem eraseLead_support_card_lt (h : f ≠ 0) : #(eraseLead f).support < #f.support := by rw [eraseLead_support] exact card_lt_card (erase_ssubset <| natDegree_mem_support_of_nonzero h) -theorem card_support_eraseLead_add_one (h : f ≠ 0) : - f.eraseLead.support.card + 1 = f.support.card := by - set c := f.support.card with hc +theorem card_support_eraseLead_add_one (h : f ≠ 0) : #f.eraseLead.support + 1 = #f.support := by + set c := #f.support with hc cases h₁ : c case zero => by_contra @@ -105,20 +104,20 @@ theorem card_support_eraseLead_add_one (h : f ≠ 0) : rfl @[simp] -theorem card_support_eraseLead : f.eraseLead.support.card = f.support.card - 1 := by +theorem card_support_eraseLead : #f.eraseLead.support = #f.support - 1 := by by_cases hf : f = 0 · rw [hf, eraseLead_zero, support_zero, card_empty] · rw [← card_support_eraseLead_add_one hf, add_tsub_cancel_right] -theorem card_support_eraseLead' {c : ℕ} (fc : f.support.card = c + 1) : - f.eraseLead.support.card = c := by +theorem card_support_eraseLead' {c : ℕ} (fc : #f.support = c + 1) : + #f.eraseLead.support = c := by rw [card_support_eraseLead, fc, add_tsub_cancel_right] theorem card_support_eq_one_of_eraseLead_eq_zero (h₀ : f ≠ 0) (h₁ : f.eraseLead = 0) : - f.support.card = 1 := + #f.support = 1 := (card_support_eq_zero.mpr h₁ ▸ card_support_eraseLead_add_one h₀).symm -theorem card_support_le_one_of_eraseLead_eq_zero (h : f.eraseLead = 0) : f.support.card ≤ 1 := by +theorem card_support_le_one_of_eraseLead_eq_zero (h : f.eraseLead = 0) : #f.support ≤ 1 := by by_cases hpz : f = 0 case pos => simp [hpz] case neg => exact le_of_eq (card_support_eq_one_of_eraseLead_eq_zero hpz h) @@ -187,7 +186,7 @@ theorem degree_eraseLead_lt (hf : f ≠ 0) : (eraseLead f).degree < f.degree := theorem eraseLead_natDegree_le_aux : (eraseLead f).natDegree ≤ f.natDegree := natDegree_le_natDegree eraseLead_degree_le -theorem eraseLead_natDegree_lt (f0 : 2 ≤ f.support.card) : (eraseLead f).natDegree < f.natDegree := +theorem eraseLead_natDegree_lt (f0 : 2 ≤ #f.support) : (eraseLead f).natDegree < f.natDegree := lt_of_le_of_ne eraseLead_natDegree_le_aux <| ne_natDegree_of_mem_eraseLead_support <| natDegree_mem_support_of_nonzero <| eraseLead_ne_zero f0 @@ -199,7 +198,7 @@ theorem natDegree_pos_of_eraseLead_ne_zero (h : f.eraseLead ≠ 0) : 0 < f.natDe theorem eraseLead_natDegree_lt_or_eraseLead_eq_zero (f : R[X]) : (eraseLead f).natDegree < f.natDegree ∨ f.eraseLead = 0 := by - by_cases h : f.support.card ≤ 1 + by_cases h : #f.support ≤ 1 · right rw [← C_mul_X_pow_eq_self h] simp @@ -264,7 +263,7 @@ theorem induction_with_natDegree_le (P : R[X] → Prop) (N : ℕ) (P_0 : P 0) (P_C_add : ∀ f g : R[X], f.natDegree < g.natDegree → g.natDegree ≤ N → P f → P g → P (f + g)) : ∀ f : R[X], f.natDegree ≤ N → P f := by intro f df - generalize hd : card f.support = c + generalize hd : #f.support = c revert f induction' c with c hc · intro f _ f0 @@ -332,7 +331,7 @@ theorem map_natDegree_eq_natDegree {S F : Type*} [Semiring S] p.natDegree.sub_zero theorem card_support_eq' {n : ℕ} (k : Fin n → ℕ) (x : Fin n → R) (hk : Function.Injective k) - (hx : ∀ i, x i ≠ 0) : (∑ i, C (x i) * X ^ k i).support.card = n := by + (hx : ∀ i, x i ≠ 0) : #(∑ i, C (x i) * X ^ k i).support = n := by suffices (∑ i, C (x i) * X ^ k i).support = image k univ by rw [this, univ.card_image_of_injective hk, card_fin] simp_rw [Finset.ext_iff, mem_support_iff, finset_sum_coeff, coeff_C_mul_X_pow, mem_image, @@ -346,7 +345,7 @@ theorem card_support_eq' {n : ℕ} (k : Fin n → ℕ) (x : Fin n → R) (hk : F · exact fun m _ hmj => if_neg fun h => hmj.symm (hk h) theorem card_support_eq {n : ℕ} : - f.support.card = n ↔ + #f.support = n ↔ ∃ (k : Fin n → ℕ) (x : Fin n → R) (_ : StrictMono k) (_ : ∀ i, x i ≠ 0), f = ∑ i, C (x i) * X ^ k i := by refine ⟨?_, fun ⟨k, x, hk, hx, hf⟩ => hf.symm ▸ card_support_eq' k x hk.injective hx⟩ @@ -391,7 +390,7 @@ theorem card_support_eq {n : ℕ} : rw [← hf, Function.extend_apply', Function.extend_apply', eraseLead_add_C_mul_X_pow] all_goals exact H -theorem card_support_eq_one : f.support.card = 1 ↔ +theorem card_support_eq_one : #f.support = 1 ↔ ∃ (k : ℕ) (x : R) (_ : x ≠ 0), f = C x * X ^ k := by refine ⟨fun h => ?_, ?_⟩ · obtain ⟨k, x, _, hx, rfl⟩ := card_support_eq.mp h @@ -400,7 +399,7 @@ theorem card_support_eq_one : f.support.card = 1 ↔ rw [support_C_mul_X_pow k hx, card_singleton] theorem card_support_eq_two : - f.support.card = 2 ↔ + #f.support = 2 ↔ ∃ (k m : ℕ) (_ : k < m) (x y : R) (_ : x ≠ 0) (_ : y ≠ 0), f = C x * X ^ k + C y * X ^ m := by refine ⟨fun h => ?_, ?_⟩ @@ -412,7 +411,7 @@ theorem card_support_eq_two : exact card_support_binomial hkm.ne hx hy theorem card_support_eq_three : - f.support.card = 3 ↔ + #f.support = 3 ↔ ∃ (k m n : ℕ) (_ : k < m) (_ : m < n) (x y z : R) (_ : x ≠ 0) (_ : y ≠ 0) (_ : z ≠ 0), f = C x * X ^ k + C y * X ^ m + C z * X ^ n := by refine ⟨fun h => ?_, ?_⟩ diff --git a/Mathlib/Algebra/Polynomial/Reverse.lean b/Mathlib/Algebra/Polynomial/Reverse.lean index 2dba679853097..614061ef26c6d 100644 --- a/Mathlib/Algebra/Polynomial/Reverse.lean +++ b/Mathlib/Algebra/Polynomial/Reverse.lean @@ -149,8 +149,8 @@ theorem reflect_monomial (N n : ℕ) : reflect N ((X : R[X]) ^ n) = X ^ revAt N theorem reflect_mul_induction (cf cg : ℕ) : ∀ N O : ℕ, ∀ f g : R[X], - f.support.card ≤ cf.succ → - g.support.card ≤ cg.succ → + #f.support ≤ cf.succ → + #g.support ≤ cg.succ → f.natDegree ≤ N → g.natDegree ≤ O → reflect (N + O) (f * g) = reflect N f * reflect O g := by induction' cf with cf hcf diff --git a/Mathlib/Algebra/Polynomial/Roots.lean b/Mathlib/Algebra/Polynomial/Roots.lean index 2c831cb0552f5..4fed6115261d6 100644 --- a/Mathlib/Algebra/Polynomial/Roots.lean +++ b/Mathlib/Algebra/Polynomial/Roots.lean @@ -27,6 +27,8 @@ We define the multiset of roots of a polynomial, and prove basic results about i -/ +open Multiset Finset + noncomputable section namespace Polynomial @@ -41,8 +43,6 @@ variable [CommRing R] [IsDomain R] {p q : R[X]} section Roots -open Multiset Finset - /-- `roots p` noncomputably gives a multiset containing all the roots of `p`, including their multiplicities. -/ noncomputable def roots (p : R[X]) : Multiset R := @@ -117,7 +117,7 @@ lemma mem_roots_iff_aeval_eq_zero {x : R} (w : p ≠ 0) : x ∈ roots p ↔ aeva Algebra.id.map_eq_id, map_id] theorem card_le_degree_of_subset_roots {p : R[X]} {Z : Finset R} (h : Z.val ⊆ p.roots) : - Z.card ≤ p.natDegree := + #Z ≤ p.natDegree := (Multiset.card_le_card (Finset.val_le_iff_val_subset.2 h)).trans (Polynomial.card_roots' p) theorem finite_setOf_isRoot {p : R[X]} (hp : p ≠ 0) : Set.Finite { x | IsRoot p x } := by @@ -548,14 +548,14 @@ lemma eq_zero_of_natDegree_lt_card_of_eval_eq_zero {R} [CommRing R] [IsDomain R] (heval : ∀ i, p.eval (f i) = 0) (hcard : natDegree p < Fintype.card ι) : p = 0 := by classical by_contra hp - apply not_lt_of_le (le_refl (Finset.card p.roots.toFinset)) + refine lt_irrefl #p.roots.toFinset ?_ calc - Finset.card p.roots.toFinset ≤ Multiset.card p.roots := Multiset.toFinset_card_le _ + #p.roots.toFinset ≤ Multiset.card p.roots := Multiset.toFinset_card_le _ _ ≤ natDegree p := Polynomial.card_roots' p _ < Fintype.card ι := hcard _ = Fintype.card (Set.range f) := (Set.card_range_of_injective hf).symm - _ = Finset.card (Finset.univ.image f) := by rw [← Set.toFinset_card, Set.toFinset_range] - _ ≤ Finset.card p.roots.toFinset := Finset.card_mono ?_ + _ = #(Finset.univ.image f) := by rw [← Set.toFinset_card, Set.toFinset_range] + _ ≤ #p.roots.toFinset := Finset.card_mono ?_ intro _ simp only [Finset.mem_image, Finset.mem_univ, true_and, Multiset.mem_toFinset, mem_roots', ne_eq, IsRoot.def, forall_exists_index, hp, not_false_eq_true] @@ -563,7 +563,7 @@ lemma eq_zero_of_natDegree_lt_card_of_eval_eq_zero {R} [CommRing R] [IsDomain R] exact heval _ lemma eq_zero_of_natDegree_lt_card_of_eval_eq_zero' {R} [CommRing R] [IsDomain R] - (p : R[X]) (s : Finset R) (heval : ∀ i ∈ s, p.eval i = 0) (hcard : natDegree p < s.card) : + (p : R[X]) (s : Finset R) (heval : ∀ i ∈ s, p.eval i = 0) (hcard : natDegree p < #s) : p = 0 := eq_zero_of_natDegree_lt_card_of_eval_eq_zero p Subtype.val_injective (fun i : s ↦ heval i i.prop) (hcard.trans_eq (Fintype.card_coe s).symm) diff --git a/Mathlib/Algebra/Polynomial/UnitTrinomial.lean b/Mathlib/Algebra/Polynomial/UnitTrinomial.lean index c21059979dbcf..10f8454bbd41f 100644 --- a/Mathlib/Algebra/Polynomial/UnitTrinomial.lean +++ b/Mathlib/Algebra/Polynomial/UnitTrinomial.lean @@ -128,7 +128,7 @@ theorem not_isUnit (hp : p.IsUnitTrinomial) : ¬IsUnit p := by ((trinomial_natDegree hkm hmn w.ne_zero).symm.trans (natDegree_eq_of_degree_eq_some (degree_eq_zero_of_isUnit h))) -theorem card_support_eq_three (hp : p.IsUnitTrinomial) : p.support.card = 3 := by +theorem card_support_eq_three (hp : p.IsUnitTrinomial) : #p.support = 3 := by obtain ⟨k, m, n, hkm, hmn, u, v, w, rfl⟩ := hp exact card_support_trinomial hkm hmn u.ne_zero v.ne_zero w.ne_zero @@ -155,7 +155,7 @@ theorem trailingCoeff_isUnit (hp : p.IsUnitTrinomial) : IsUnit p.trailingCoeff : end IsUnitTrinomial theorem isUnitTrinomial_iff : - p.IsUnitTrinomial ↔ p.support.card = 3 ∧ ∀ k ∈ p.support, IsUnit (p.coeff k) := by + p.IsUnitTrinomial ↔ #p.support = 3 ∧ ∀ k ∈ p.support, IsUnit (p.coeff k) := by refine ⟨fun hp => ⟨hp.card_support_eq_three, fun k => hp.coeff_isUnit⟩, fun hp => ?_⟩ obtain ⟨k, m, n, hkm, hmn, x, y, z, hx, hy, hz, rfl⟩ := card_support_eq_three.mp hp.1 rw [support_trinomial hkm hmn hx hy hz] at hp diff --git a/Mathlib/Combinatorics/SimpleGraph/Regularity/Increment.lean b/Mathlib/Combinatorics/SimpleGraph/Regularity/Increment.lean index 4fe6e82a487fa..a6ff90799f9cb 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Regularity/Increment.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Regularity/Increment.lean @@ -66,7 +66,7 @@ theorem card_increment (hPα : P.parts.card * 16 ^ P.parts.card ≤ card α) (hP have hPpos : 0 < stepBound P.parts.card := stepBound_pos (nonempty_of_not_uniform hPG).card_pos rw [increment, card_bind] simp_rw [chunk, apply_dite Finpartition.parts, apply_dite card, sum_dite] - rw [sum_const_nat, sum_const_nat, card_attach, card_attach]; rotate_left + rw [sum_const_nat, sum_const_nat, univ_eq_attach, univ_eq_attach, card_attach, card_attach] any_goals exact fun x hx => card_parts_equitabilise _ _ (Nat.div_pos hPα' hPpos).ne' rw [Nat.sub_add_cancel a_add_one_le_four_pow_parts_card, Nat.sub_add_cancel ((Nat.le_succ _).trans a_add_one_le_four_pow_parts_card), ← add_mul] From 818db408b33f801f4094614250e6b3cfb9e065c8 Mon Sep 17 00:00:00 2001 From: damiano Date: Tue, 22 Oct 2024 13:18:33 +0000 Subject: [PATCH 271/505] fix: silence the multigoal linter with fail_... tactics (#18063) Reported on [Zulip](https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/false.20positive.20of.20linter.2Estyle.2EmultiGoal) --- Mathlib/Tactic/Linter/Multigoal.lean | 4 +++- test/Multigoal.lean | 8 ++++++++ 2 files changed, 11 insertions(+), 1 deletion(-) diff --git a/Mathlib/Tactic/Linter/Multigoal.lean b/Mathlib/Tactic/Linter/Multigoal.lean index 5502f129e1887..aa24f8a50aef1 100644 --- a/Mathlib/Tactic/Linter/Multigoal.lean +++ b/Mathlib/Tactic/Linter/Multigoal.lean @@ -118,7 +118,9 @@ abbrev ignoreBranch : Std.HashSet SyntaxNodeKind := .ofList [ ``Lean.Parser.Tactic.tacticIterate____, ``Lean.Parser.Tactic.anyGoals, ``Lean.Parser.Tactic.allGoals, - ``Lean.Parser.Tactic.focus + ``Lean.Parser.Tactic.focus, + ``Lean.Parser.Tactic.failIfSuccess, + `Mathlib.Tactic.successIfFailWithMsg ] /-- `getManyGoals t` returns the syntax nodes of the `InfoTree` `t` corresponding to tactic calls diff --git a/test/Multigoal.lean b/test/Multigoal.lean index a18e6bf88aad5..e43848c0295dc 100644 --- a/test/Multigoal.lean +++ b/test/Multigoal.lean @@ -2,6 +2,7 @@ import Mathlib.Tactic.Basic import Mathlib.Tactic.Conv import Mathlib.Tactic.Linter.Multigoal import Mathlib.Util.SleepHeartbeats +import Mathlib.Tactic.SuccessIfFailWithMsg -- The warning generated by `linter.style.multiGoal` is not suppressed by `#guard_msgs`, -- because the linter is run on `#guard_msgs` itself. This is a known issue, see e.g. @@ -121,3 +122,10 @@ set_option linter.style.multiGoal true in example : True ∧ True := by constructor bi_trivial + +-- Exclude `fail_if_success` and `success_if_fail_with_msg` from linting. +set_option linter.style.multiGoal true in +example : True := by + fail_if_success done + success_if_fail_with_msg "internal exception #4" done + exact .intro From ac12c8737d7642ab60f988f0dad06c645b6e3180 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 22 Oct 2024 13:55:26 +0000 Subject: [PATCH 272/505] chore(Algebra/Polynomial/RingDivision): move lemmas earlier (#18050) `Algebra.Polynomial.RingDivision` is a bag grab of lemmas with absolutely no running theme. This PR moves two thirds of them to earlier files with no proof change. --- Mathlib/Algebra/Polynomial/AlgebraMap.lean | 72 ++- Mathlib/Algebra/Polynomial/BigOperators.lean | 10 + .../Polynomial/Degree/Definitions.lean | 238 ++++++-- Mathlib/Algebra/Polynomial/Degree/Lemmas.lean | 25 + Mathlib/Algebra/Polynomial/Div.lean | 173 +++++- Mathlib/Algebra/Polynomial/Eval.lean | 54 ++ Mathlib/Algebra/Polynomial/Mirror.lean | 2 +- Mathlib/Algebra/Polynomial/Monic.lean | 88 +++ Mathlib/Algebra/Polynomial/RingDivision.lean | 565 +----------------- Mathlib/Algebra/Polynomial/Roots.lean | 2 +- .../Polynomial/SumIteratedDerivative.lean | 1 - Mathlib/Algebra/Polynomial/UnitTrinomial.lean | 1 + Mathlib/Analysis/SpecificLimits/Normed.lean | 1 + Mathlib/FieldTheory/RatFunc/Defs.lean | 2 +- .../Polynomial/IrreducibleRing.lean | 1 - Mathlib/RingTheory/Polynomial/Pochhammer.lean | 3 +- scripts/no_lints_prime_decls.txt | 1 - 17 files changed, 613 insertions(+), 626 deletions(-) diff --git a/Mathlib/Algebra/Polynomial/AlgebraMap.lean b/Mathlib/Algebra/Polynomial/AlgebraMap.lean index 29c855d09a69b..1f0fbd0dedb84 100644 --- a/Mathlib/Algebra/Polynomial/AlgebraMap.lean +++ b/Mathlib/Algebra/Polynomial/AlgebraMap.lean @@ -531,11 +531,30 @@ theorem not_isUnit_X_sub_C [Nontrivial R] (r : R) : ¬IsUnit (X - C r) := end Ring -theorem aeval_endomorphism {M : Type*} [CommRing R] [AddCommGroup M] [Module R M] (f : M →ₗ[R] M) +section CommRing +variable [CommRing R] {p : R[X]} {t : R} + +theorem aeval_endomorphism {M : Type*} [AddCommGroup M] [Module R M] (f : M →ₗ[R] M) (v : M) (p : R[X]) : aeval f p v = p.sum fun n b => b • (f ^ n) v := by rw [aeval_def, eval₂_eq_sum] exact map_sum (LinearMap.applyₗ v) _ _ +lemma X_sub_C_pow_dvd_iff {n : ℕ} : (X - C t) ^ n ∣ p ↔ X ^ n ∣ p.comp (X + C t) := by + convert (map_dvd_iff <| algEquivAevalXAddC t).symm using 2 + simp [C_eq_algebraMap] + +lemma comp_X_add_C_eq_zero_iff : p.comp (X + C t) = 0 ↔ p = 0 := + AddEquivClass.map_eq_zero_iff (algEquivAevalXAddC t) + +lemma comp_X_add_C_ne_zero_iff : p.comp (X + C t) ≠ 0 ↔ p ≠ 0 := comp_X_add_C_eq_zero_iff.not + +variable [IsDomain R] + +lemma units_coeff_zero_smul (c : R[X]ˣ) (p : R[X]) : (c : R[X]).coeff 0 • p = c * p := by + rw [← Polynomial.C_mul', ← Polynomial.eq_C_of_degree_eq_zero (degree_coe_units c)] + +end CommRing + section StableSubmodule variable {M : Type*} [CommSemiring R] [AddCommMonoid M] [Module R M] @@ -561,4 +580,55 @@ lemma aeval_apply_smul_mem_of_le_comap end StableSubmodule +section CommSemiring + +variable [CommSemiring R] {a p : R[X]} + +theorem eq_zero_of_mul_eq_zero_of_smul (P : R[X]) (h : ∀ r : R, r • P = 0 → r = 0) : + ∀ (Q : R[X]), P * Q = 0 → Q = 0 := by + intro Q hQ + suffices ∀ i, P.coeff i • Q = 0 by + rw [← leadingCoeff_eq_zero] + apply h + simpa [ext_iff, mul_comm Q.leadingCoeff] using fun i ↦ congr_arg (·.coeff Q.natDegree) (this i) + apply Nat.strong_decreasing_induction + · use P.natDegree + intro i hi + rw [coeff_eq_zero_of_natDegree_lt hi, zero_smul] + intro l IH + obtain _|hl := (natDegree_smul_le (P.coeff l) Q).lt_or_eq + · apply eq_zero_of_mul_eq_zero_of_smul _ h (P.coeff l • Q) + rw [smul_eq_C_mul, mul_left_comm, hQ, mul_zero] + suffices P.coeff l * Q.leadingCoeff = 0 by + rwa [← leadingCoeff_eq_zero, ← coeff_natDegree, coeff_smul, hl, coeff_natDegree, smul_eq_mul] + let m := Q.natDegree + suffices (P * Q).coeff (l + m) = P.coeff l * Q.leadingCoeff by rw [← this, hQ, coeff_zero] + rw [coeff_mul] + apply Finset.sum_eq_single (l, m) _ (by simp) + simp only [Finset.mem_antidiagonal, ne_eq, Prod.forall, Prod.mk.injEq, not_and] + intro i j hij H + obtain hi|rfl|hi := lt_trichotomy i l + · have hj : m < j := by omega + rw [coeff_eq_zero_of_natDegree_lt hj, mul_zero] + · omega + · rw [← coeff_C_mul, ← smul_eq_C_mul, IH _ hi, coeff_zero] +termination_by Q => Q.natDegree + +open nonZeroDivisors in +/-- *McCoy theorem*: a polynomial `P : R[X]` is a zerodivisor if and only if there is `a : R` +such that `a ≠ 0` and `a • P = 0`. -/ +theorem nmem_nonZeroDivisors_iff {P : R[X]} : P ∉ R[X]⁰ ↔ ∃ a : R, a ≠ 0 ∧ a • P = 0 := by + refine ⟨fun hP ↦ ?_, fun ⟨a, ha, h⟩ h1 ↦ ha <| C_eq_zero.1 <| (h1 _) <| smul_eq_C_mul a ▸ h⟩ + by_contra! h + obtain ⟨Q, hQ⟩ := _root_.nmem_nonZeroDivisors_iff.1 hP + refine hQ.2 (eq_zero_of_mul_eq_zero_of_smul P (fun a ha ↦ ?_) Q (mul_comm P _ ▸ hQ.1)) + contrapose! ha + exact h a ha + +open nonZeroDivisors in +protected lemma mem_nonZeroDivisors_iff {P : R[X]} : P ∈ R[X]⁰ ↔ ∀ a : R, a • P = 0 → a = 0 := by + simpa [not_imp_not] using (nmem_nonZeroDivisors_iff (P := P)).not + +end CommSemiring + end Polynomial diff --git a/Mathlib/Algebra/Polynomial/BigOperators.lean b/Mathlib/Algebra/Polynomial/BigOperators.lean index 53e09909c6d90..557f139b15527 100644 --- a/Mathlib/Algebra/Polynomial/BigOperators.lean +++ b/Mathlib/Algebra/Polynomial/BigOperators.lean @@ -270,6 +270,16 @@ theorem prod_X_sub_C_coeff_card_pred (s : Finset ι) (f : ι → R) (hs : 0 < #s (∏ i ∈ s, (X - C (f i))).coeff (#s - 1) = -∑ i ∈ s, f i := by simpa using multiset_prod_X_sub_C_coeff_card_pred (s.1.map f) (by simpa using hs) +variable [IsDomain R] + +@[simp] +lemma natDegree_multiset_prod_X_sub_C_eq_card (s : Multiset R) : + (s.map (X - C ·)).prod.natDegree = Multiset.card s := by + rw [natDegree_multiset_prod_of_monic, Multiset.map_map] + · simp only [(· ∘ ·), natDegree_X_sub_C, Multiset.map_const', Multiset.sum_replicate, smul_eq_mul, + mul_one] + · exact Multiset.forall_mem_map_iff.2 fun a _ => monic_X_sub_C a + end CommRing section NoZeroDivisors diff --git a/Mathlib/Algebra/Polynomial/Degree/Definitions.lean b/Mathlib/Algebra/Polynomial/Degree/Definitions.lean index 80fccb38b65f8..9ec69895a9634 100644 --- a/Mathlib/Algebra/Polynomial/Degree/Definitions.lean +++ b/Mathlib/Algebra/Polynomial/Degree/Definitions.lean @@ -1370,8 +1370,188 @@ theorem leadingCoeff_pow_X_add_C (r : R) (i : ℕ) : leadingCoeff ((X + C r) ^ i nontriviality rw [leadingCoeff_pow'] <;> simp +variable [NoZeroDivisors R] {p q : R[X]} + +@[simp] +lemma degree_mul : degree (p * q) = degree p + degree q := + letI := Classical.decEq R + if hp0 : p = 0 then by simp only [hp0, degree_zero, zero_mul, WithBot.bot_add] + else + if hq0 : q = 0 then by simp only [hq0, degree_zero, mul_zero, WithBot.add_bot] + else degree_mul' <| mul_ne_zero (mt leadingCoeff_eq_zero.1 hp0) (mt leadingCoeff_eq_zero.1 hq0) + +/-- `degree` as a monoid homomorphism between `R[X]` and `Multiplicative (WithBot ℕ)`. + This is useful to prove results about multiplication and degree. -/ +def degreeMonoidHom [Nontrivial R] : R[X] →* Multiplicative (WithBot ℕ) where + toFun := degree + map_one' := degree_one + map_mul' _ _ := degree_mul + +@[simp] +lemma degree_pow [Nontrivial R] (p : R[X]) (n : ℕ) : degree (p ^ n) = n • degree p := + map_pow (@degreeMonoidHom R _ _ _) _ _ + +@[simp] +lemma leadingCoeff_mul (p q : R[X]) : leadingCoeff (p * q) = leadingCoeff p * leadingCoeff q := by + by_cases hp : p = 0 + · simp only [hp, zero_mul, leadingCoeff_zero] + · by_cases hq : q = 0 + · simp only [hq, mul_zero, leadingCoeff_zero] + · rw [leadingCoeff_mul'] + exact mul_ne_zero (mt leadingCoeff_eq_zero.1 hp) (mt leadingCoeff_eq_zero.1 hq) + +/-- `Polynomial.leadingCoeff` bundled as a `MonoidHom` when `R` has `NoZeroDivisors`, and thus + `leadingCoeff` is multiplicative -/ +def leadingCoeffHom : R[X] →* R where + toFun := leadingCoeff + map_one' := by simp + map_mul' := leadingCoeff_mul + +@[simp] +lemma leadingCoeffHom_apply (p : R[X]) : leadingCoeffHom p = leadingCoeff p := + rfl + +@[simp] +lemma leadingCoeff_pow (p : R[X]) (n : ℕ) : leadingCoeff (p ^ n) = leadingCoeff p ^ n := + (leadingCoeffHom : R[X] →* R).map_pow p n + +lemma leadingCoeff_dvd_leadingCoeff {a p : R[X]} (hap : a ∣ p) : + a.leadingCoeff ∣ p.leadingCoeff := + map_dvd leadingCoeffHom hap + +instance : NoZeroDivisors R[X] where + eq_zero_or_eq_zero_of_mul_eq_zero h := by + rw [← leadingCoeff_eq_zero, ← leadingCoeff_eq_zero] + refine eq_zero_or_eq_zero_of_mul_eq_zero ?_ + rw [← leadingCoeff_zero, ← leadingCoeff_mul, h] +lemma natDegree_mul (hp : p ≠ 0) (hq : q ≠ 0) : (p*q).natDegree = p.natDegree + q.natDegree := by + rw [← Nat.cast_inj (R := WithBot ℕ), ← degree_eq_natDegree (mul_ne_zero hp hq), + Nat.cast_add, ← degree_eq_natDegree hp, ← degree_eq_natDegree hq, degree_mul] + +@[simp] +lemma natDegree_pow (p : R[X]) (n : ℕ) : natDegree (p ^ n) = n * natDegree p := by + classical + obtain rfl | hp := eq_or_ne p 0 + · obtain rfl | hn := eq_or_ne n 0 <;> simp [*] + exact natDegree_pow' <| by + rw [← leadingCoeff_pow, Ne, leadingCoeff_eq_zero]; exact pow_ne_zero _ hp + +lemma degree_le_mul_left (p : R[X]) (hq : q ≠ 0) : degree p ≤ degree (p * q) := by + classical + obtain rfl | hp := eq_or_ne p 0 + · simp + · rw [degree_mul, degree_eq_natDegree hp, degree_eq_natDegree hq] + exact WithBot.coe_le_coe.2 (Nat.le_add_right _ _) + +lemma natDegree_le_of_dvd (h1 : p ∣ q) (h2 : q ≠ 0) : p.natDegree ≤ q.natDegree := by + obtain ⟨q, rfl⟩ := h1 + rw [mul_ne_zero_iff] at h2 + rw [natDegree_mul h2.1 h2.2]; exact Nat.le_add_right _ _ + +lemma degree_le_of_dvd (h1 : p ∣ q) (h2 : q ≠ 0) : degree p ≤ degree q := by + rcases h1 with ⟨q, rfl⟩; rw [mul_ne_zero_iff] at h2 + exact degree_le_mul_left p h2.2 + +lemma eq_zero_of_dvd_of_degree_lt (h₁ : p ∣ q) (h₂ : degree q < degree p) : q = 0 := by + by_contra hc + exact (lt_iff_not_ge _ _).mp h₂ (degree_le_of_dvd h₁ hc) + +lemma eq_zero_of_dvd_of_natDegree_lt (h₁ : p ∣ q) (h₂ : natDegree q < natDegree p) : + q = 0 := by + by_contra hc + exact (lt_iff_not_ge _ _).mp h₂ (natDegree_le_of_dvd h₁ hc) + +lemma not_dvd_of_degree_lt (h0 : q ≠ 0) (hl : q.degree < p.degree) : ¬p ∣ q := by + by_contra hcontra + exact h0 (eq_zero_of_dvd_of_degree_lt hcontra hl) + +lemma not_dvd_of_natDegree_lt (h0 : q ≠ 0) (hl : q.natDegree < p.natDegree) : + ¬p ∣ q := by + by_contra hcontra + exact h0 (eq_zero_of_dvd_of_natDegree_lt hcontra hl) + +/-- This lemma is useful for working with the `intDegree` of a rational function. -/ +lemma natDegree_sub_eq_of_prod_eq {p₁ p₂ q₁ q₂ : R[X]} (hp₁ : p₁ ≠ 0) (hq₁ : q₁ ≠ 0) + (hp₂ : p₂ ≠ 0) (hq₂ : q₂ ≠ 0) (h_eq : p₁ * q₂ = p₂ * q₁) : + (p₁.natDegree : ℤ) - q₁.natDegree = (p₂.natDegree : ℤ) - q₂.natDegree := by + rw [sub_eq_sub_iff_add_eq_add] + norm_cast + rw [← natDegree_mul hp₁ hq₂, ← natDegree_mul hp₂ hq₁, h_eq] + +lemma natDegree_eq_zero_of_isUnit (h : IsUnit p) : natDegree p = 0 := by + nontriviality R + obtain ⟨q, hq⟩ := h.exists_right_inv + have := natDegree_mul (left_ne_zero_of_mul_eq_one hq) (right_ne_zero_of_mul_eq_one hq) + rw [hq, natDegree_one, eq_comm, add_eq_zero] at this + exact this.1 + +lemma degree_eq_zero_of_isUnit [Nontrivial R] (h : IsUnit p) : degree p = 0 := + (natDegree_eq_zero_iff_degree_le_zero.mp <| natDegree_eq_zero_of_isUnit h).antisymm + (zero_le_degree_iff.mpr h.ne_zero) + +@[simp] +lemma degree_coe_units [Nontrivial R] (u : R[X]ˣ) : degree (u : R[X]) = 0 := + degree_eq_zero_of_isUnit ⟨u, rfl⟩ + +/-- Characterization of a unit of a polynomial ring over an integral domain `R`. +See `Polynomial.isUnit_iff_coeff_isUnit_isNilpotent` when `R` is a commutative ring. -/ +lemma isUnit_iff : IsUnit p ↔ ∃ r : R, IsUnit r ∧ C r = p := + ⟨fun hp => + ⟨p.coeff 0, + let h := eq_C_of_natDegree_eq_zero (natDegree_eq_zero_of_isUnit hp) + ⟨isUnit_C.1 (h ▸ hp), h.symm⟩⟩, + fun ⟨_, hr, hrp⟩ => hrp ▸ isUnit_C.2 hr⟩ + +lemma not_isUnit_of_degree_pos (p : R[X]) (hpl : 0 < p.degree) : ¬ IsUnit p := by + cases subsingleton_or_nontrivial R + · simp [Subsingleton.elim p 0] at hpl + intro h + simp [degree_eq_zero_of_isUnit h] at hpl + +lemma not_isUnit_of_natDegree_pos (p : R[X]) (hpl : 0 < p.natDegree) : ¬ IsUnit p := + not_isUnit_of_degree_pos _ (natDegree_pos_iff_degree_pos.mp hpl) + +@[simp] lemma natDegree_coe_units (u : R[X]ˣ) : natDegree (u : R[X]) = 0 := by + nontriviality R + exact natDegree_eq_of_degree_eq_some (degree_coe_units u) + +theorem coeff_coe_units_zero_ne_zero [Nontrivial R] (u : R[X]ˣ) : coeff (u : R[X]) 0 ≠ 0 := by + conv in 0 => rw [← natDegree_coe_units u] + rw [← leadingCoeff, Ne, leadingCoeff_eq_zero] + exact Units.ne_zero _ + end Semiring +section CommSemiring +variable [CommSemiring R] {a p : R[X]} (hp : p.Monic) +include hp + +lemma Monic.C_dvd_iff_isUnit {a : R} : C a ∣ p ↔ IsUnit a where + mp h := isUnit_iff_dvd_one.mpr <| hp.coeff_natDegree ▸ (C_dvd_iff_dvd_coeff _ _).mp h p.natDegree + mpr ha := (ha.map C).dvd + +lemma Monic.natDegree_pos : 0 < natDegree p ↔ p ≠ 1 := + Nat.pos_iff_ne_zero.trans hp.natDegree_eq_zero.not + +lemma Monic.degree_pos : 0 < degree p ↔ p ≠ 1 := + natDegree_pos_iff_degree_pos.symm.trans hp.natDegree_pos + +lemma Monic.degree_pos_of_not_isUnit (hu : ¬IsUnit p) : 0 < degree p := + hp.degree_pos.mpr fun hp' ↦ (hp' ▸ hu) isUnit_one + +lemma Monic.natDegree_pos_of_not_isUnit (hu : ¬IsUnit p) : 0 < natDegree p := + hp.natDegree_pos.mpr fun hp' ↦ (hp' ▸ hu) isUnit_one + +lemma degree_pos_of_not_isUnit_of_dvd_monic (ha : ¬IsUnit a) (hap : a ∣ p) : 0 < degree a := by + contrapose! ha with h + rw [Polynomial.eq_C_of_degree_le_zero h] at hap ⊢ + simpa [hp.C_dvd_iff_isUnit, isUnit_C] using hap + +lemma natDegree_pos_of_not_isUnit_of_dvd_monic (ha : ¬IsUnit a) (hap : a ∣ p) : 0 < natDegree a := + natDegree_pos_iff_degree_pos.mpr <| degree_pos_of_not_isUnit_of_dvd_monic hp ha hap + +end CommSemiring + section Ring variable [Ring R] @@ -1420,59 +1600,11 @@ theorem natDegree_X_pow_sub_C {n : ℕ} {r : R} : (X ^ n - C r).natDegree = n := theorem leadingCoeff_X_sub_C [Ring S] (r : S) : (X - C r).leadingCoeff = 1 := by rw [sub_eq_add_neg, ← map_neg C r, leadingCoeff_X_add_C] -end Ring - -section NoZeroDivisors - -variable [Semiring R] [NoZeroDivisors R] {p q : R[X]} - -@[simp] -theorem degree_mul : degree (p * q) = degree p + degree q := - letI := Classical.decEq R - if hp0 : p = 0 then by simp only [hp0, degree_zero, zero_mul, WithBot.bot_add] - else - if hq0 : q = 0 then by simp only [hq0, degree_zero, mul_zero, WithBot.add_bot] - else degree_mul' <| mul_ne_zero (mt leadingCoeff_eq_zero.1 hp0) (mt leadingCoeff_eq_zero.1 hq0) - -/-- `degree` as a monoid homomorphism between `R[X]` and `Multiplicative (WithBot ℕ)`. - This is useful to prove results about multiplication and degree. -/ -def degreeMonoidHom [Nontrivial R] : R[X] →* Multiplicative (WithBot ℕ) where - toFun := degree - map_one' := degree_one - map_mul' _ _ := degree_mul +variable [IsDomain R] {p q : R[X]} -@[simp] -theorem degree_pow [Nontrivial R] (p : R[X]) (n : ℕ) : degree (p ^ n) = n • degree p := - map_pow (@degreeMonoidHom R _ _ _) _ _ - -@[simp] -theorem leadingCoeff_mul (p q : R[X]) : leadingCoeff (p * q) = leadingCoeff p * leadingCoeff q := by - by_cases hp : p = 0 - · simp only [hp, zero_mul, leadingCoeff_zero] - · by_cases hq : q = 0 - · simp only [hq, mul_zero, leadingCoeff_zero] - · rw [leadingCoeff_mul'] - exact mul_ne_zero (mt leadingCoeff_eq_zero.1 hp) (mt leadingCoeff_eq_zero.1 hq) - -/-- `Polynomial.leadingCoeff` bundled as a `MonoidHom` when `R` has `NoZeroDivisors`, and thus - `leadingCoeff` is multiplicative -/ -def leadingCoeffHom : R[X] →* R where - toFun := leadingCoeff - map_one' := by simp - map_mul' := leadingCoeff_mul - -@[simp] -theorem leadingCoeffHom_apply (p : R[X]) : leadingCoeffHom p = leadingCoeff p := - rfl - -@[simp] -theorem leadingCoeff_pow (p : R[X]) (n : ℕ) : leadingCoeff (p ^ n) = leadingCoeff p ^ n := - (leadingCoeffHom : R[X] →* R).map_pow p n - -theorem leadingCoeff_dvd_leadingCoeff {a p : R[X]} (hap : a ∣ p) : - a.leadingCoeff ∣ p.leadingCoeff := - map_dvd leadingCoeffHom hap - -end NoZeroDivisors +instance : IsDomain R[X] := NoZeroDivisors.to_isDomain _ +end Ring end Polynomial + +set_option linter.style.longFile 1700 diff --git a/Mathlib/Algebra/Polynomial/Degree/Lemmas.lean b/Mathlib/Algebra/Polynomial/Degree/Lemmas.lean index 649bab3a13526..00c456c6a6f64 100644 --- a/Mathlib/Algebra/Polynomial/Degree/Lemmas.lean +++ b/Mathlib/Algebra/Polynomial/Degree/Lemmas.lean @@ -368,6 +368,31 @@ theorem leadingCoeff_comp (hq : natDegree q ≠ 0) : end NoZeroDivisors +section CommRing +variable [CommRing R] {p q : R[X]} + +@[simp] lemma 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] + +variable [IsDomain R] + +lemma comp_eq_zero_iff : p.comp q = 0 ↔ p = 0 ∨ p.eval (q.coeff 0) = 0 ∧ q = C (q.coeff 0) := by + refine ⟨fun h ↦ ?_, Or.rec (fun h ↦ by simp [h]) fun h ↦ by rw [h.2, comp_C, h.1, C_0]⟩ + have key : p.natDegree = 0 ∨ q.natDegree = 0 := by + rw [← mul_eq_zero, ← natDegree_comp, h, natDegree_zero] + obtain key | key := Or.imp eq_C_of_natDegree_eq_zero eq_C_of_natDegree_eq_zero key + · rw [key, C_comp] at h + exact Or.inl (key.trans h) + · rw [key, comp_C, C_eq_zero] at h + exact Or.inr ⟨h, key⟩ + +end CommRing + section DivisionRing variable {K : Type*} [DivisionRing K] diff --git a/Mathlib/Algebra/Polynomial/Div.lean b/Mathlib/Algebra/Polynomial/Div.lean index 50f666e013f11..804749c1a188a 100644 --- a/Mathlib/Algebra/Polynomial/Div.lean +++ b/Mathlib/Algebra/Polynomial/Div.lean @@ -556,7 +556,7 @@ end Ring section CommRing -variable [CommRing R] {p q : R[X]} +variable [CommRing R] {p p₁ p₂ q : R[X]} @[simp] theorem modByMonic_X_sub_C_eq_C_eval (p : R[X]) (a : R) : p %ₘ (X - C a) = C (p.eval a) := by @@ -659,6 +659,177 @@ lemma mul_self_modByMonic (hq : q.Monic) : (p * q) %ₘ q = 0 := by rw [modByMonic_eq_zero_iff_dvd hq] exact dvd_mul_left q p +lemma modByMonic_eq_of_dvd_sub (hq : q.Monic) (h : q ∣ p₁ - p₂) : p₁ %ₘ q = p₂ %ₘ q := by + nontriviality R + obtain ⟨f, sub_eq⟩ := h + refine (div_modByMonic_unique (p₂ /ₘ q + f) _ hq ⟨?_, degree_modByMonic_lt _ hq⟩).2 + rw [sub_eq_iff_eq_add.mp sub_eq, mul_add, ← add_assoc, modByMonic_add_div _ hq, add_comm] + +lemma add_modByMonic (p₁ p₂ : R[X]) : (p₁ + p₂) %ₘ q = p₁ %ₘ q + p₂ %ₘ q := by + by_cases hq : q.Monic + · cases' subsingleton_or_nontrivial R with hR hR + · simp only [eq_iff_true_of_subsingleton] + · exact + (div_modByMonic_unique (p₁ /ₘ q + p₂ /ₘ q) _ hq + ⟨by + rw [mul_add, add_left_comm, add_assoc, modByMonic_add_div _ hq, ← add_assoc, + add_comm (q * _), modByMonic_add_div _ hq], + (degree_add_le _ _).trans_lt + (max_lt (degree_modByMonic_lt _ hq) (degree_modByMonic_lt _ hq))⟩).2 + · simp_rw [modByMonic_eq_of_not_monic _ hq] + +lemma neg_modByMonic (p q : R[X]) : (-p) %ₘ q = - (p %ₘ q) := by + rw [eq_neg_iff_add_eq_zero, ← add_modByMonic, neg_add_cancel, zero_modByMonic] + +lemma sub_modByMonic (p₁ p₂ q : R[X]) : (p₁ - p₂) %ₘ q = p₁ %ₘ q - p₂ %ₘ q := by + simp [sub_eq_add_neg, add_modByMonic, neg_modByMonic] + +lemma eval_divByMonic_eq_trailingCoeff_comp {p : R[X]} {t : R} : + (p /ₘ (X - C t) ^ p.rootMultiplicity t).eval t = (p.comp (X + C t)).trailingCoeff := by + obtain rfl | hp := eq_or_ne p 0 + · rw [zero_divByMonic, eval_zero, zero_comp, trailingCoeff_zero] + have mul_eq := p.pow_mul_divByMonic_rootMultiplicity_eq t + set m := p.rootMultiplicity t + set g := p /ₘ (X - C t) ^ m + have : (g.comp (X + C t)).coeff 0 = g.eval t := by + rw [coeff_zero_eq_eval_zero, eval_comp, eval_add, eval_X, eval_C, zero_add] + rw [← congr_arg (comp · <| X + C t) mul_eq, mul_comp, pow_comp, sub_comp, X_comp, C_comp, + add_sub_cancel_right, ← reverse_leadingCoeff, reverse_X_pow_mul, reverse_leadingCoeff, + trailingCoeff, Nat.le_zero.1 (natTrailingDegree_le_of_ne_zero <| + this ▸ eval_divByMonic_pow_rootMultiplicity_ne_zero t hp), this] + +/- Porting note: the ML3 proof no longer worked because of a conflict in the +inferred type and synthesized type for `DecidableRel` when using `Nat.le_find_iff` from +`Mathlib.Algebra.Polynomial.Div` After some discussion on [Zulip] +(https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/decidability.20leakage) +introduced `Polynomial.rootMultiplicity_eq_nat_find_of_nonzero` to contain the issue +-/ +/-- The multiplicity of `a` as root of a nonzero polynomial `p` is at least `n` iff +`(X - a) ^ n` divides `p`. -/ +lemma le_rootMultiplicity_iff (p0 : p ≠ 0) {a : R} {n : ℕ} : + n ≤ rootMultiplicity a p ↔ (X - C a) ^ n ∣ p := by + classical + rw [rootMultiplicity_eq_nat_find_of_nonzero p0, @Nat.le_find_iff _ (_)] + simp_rw [Classical.not_not] + refine ⟨fun h => ?_, fun h m hm => (pow_dvd_pow _ hm).trans h⟩ + cases' n with n + · rw [pow_zero] + apply one_dvd + · exact h n n.lt_succ_self + +lemma rootMultiplicity_le_iff (p0 : p ≠ 0) (a : R) (n : ℕ) : + rootMultiplicity a p ≤ n ↔ ¬(X - C a) ^ (n + 1) ∣ p := by + rw [← (le_rootMultiplicity_iff p0).not, not_le, Nat.lt_add_one_iff] + +/-- The multiplicity of `p + q` is at least the minimum of the multiplicities. -/ +lemma rootMultiplicity_add {p q : R[X]} (a : R) (hzero : p + q ≠ 0) : + min (rootMultiplicity a p) (rootMultiplicity a q) ≤ rootMultiplicity a (p + q) := by + rw [le_rootMultiplicity_iff hzero] + exact min_pow_dvd_add (pow_rootMultiplicity_dvd p a) (pow_rootMultiplicity_dvd q a) + +lemma le_rootMultiplicity_mul {p q : R[X]} (x : R) (hpq : p * q ≠ 0) : + rootMultiplicity x p + rootMultiplicity x q ≤ rootMultiplicity x (p * q) := by + rw [le_rootMultiplicity_iff hpq, pow_add] + exact mul_dvd_mul (pow_rootMultiplicity_dvd p x) (pow_rootMultiplicity_dvd q x) + +lemma pow_rootMultiplicity_not_dvd (p0 : p ≠ 0) (a : R) : + ¬(X - C a) ^ (rootMultiplicity a p + 1) ∣ p := by rw [← rootMultiplicity_le_iff p0] + +/-- See `Polynomial.rootMultiplicity_eq_natTrailingDegree` for the general case. -/ +lemma rootMultiplicity_eq_natTrailingDegree' : p.rootMultiplicity 0 = p.natTrailingDegree := by + by_cases h : p = 0 + · simp only [h, rootMultiplicity_zero, natTrailingDegree_zero] + refine le_antisymm ?_ ?_ + · rw [rootMultiplicity_le_iff h, map_zero, sub_zero, X_pow_dvd_iff, not_forall] + exact ⟨p.natTrailingDegree, + fun h' ↦ trailingCoeff_nonzero_iff_nonzero.2 h <| h' <| Nat.lt.base _⟩ + · rw [le_rootMultiplicity_iff h, map_zero, sub_zero, X_pow_dvd_iff] + exact fun _ ↦ coeff_eq_zero_of_lt_natTrailingDegree + +/-- Division by a monic polynomial doesn't change the leading coefficient. -/ +lemma leadingCoeff_divByMonic_of_monic (hmonic : q.Monic) + (hdegree : q.degree ≤ p.degree) : (p /ₘ q).leadingCoeff = p.leadingCoeff := by + nontriviality + have h : q.leadingCoeff * (p /ₘ q).leadingCoeff ≠ 0 := by + simpa [divByMonic_eq_zero_iff hmonic, hmonic.leadingCoeff, + Nat.WithBot.one_le_iff_zero_lt] using hdegree + nth_rw 2 [← modByMonic_add_div p hmonic] + rw [leadingCoeff_add_of_degree_lt, leadingCoeff_monic_mul hmonic] + rw [degree_mul' h, degree_add_divByMonic hmonic hdegree] + exact (degree_modByMonic_lt p hmonic).trans_le hdegree + +variable [IsDomain R] + +lemma degree_eq_one_of_irreducible_of_root (hi : Irreducible p) {x : R} (hx : IsRoot p x) : + degree p = 1 := + let ⟨g, hg⟩ := dvd_iff_isRoot.2 hx + have : IsUnit (X - C x) ∨ IsUnit g := hi.isUnit_or_isUnit hg + this.elim + (fun h => by + have h₁ : degree (X - C x) = 1 := degree_X_sub_C x + have h₂ : degree (X - C x) = 0 := degree_eq_zero_of_isUnit h + rw [h₁] at h₂; exact absurd h₂ (by decide)) + fun hgu => by rw [hg, degree_mul, degree_X_sub_C, degree_eq_zero_of_isUnit hgu, add_zero] + +lemma leadingCoeff_divByMonic_X_sub_C (p : R[X]) (hp : degree p ≠ 0) (a : R) : + leadingCoeff (p /ₘ (X - C a)) = leadingCoeff p := by + nontriviality + cases' hp.lt_or_lt with hd hd + · rw [degree_eq_bot.mp <| Nat.WithBot.lt_zero_iff.mp hd, zero_divByMonic] + refine leadingCoeff_divByMonic_of_monic (monic_X_sub_C a) ?_ + rwa [degree_X_sub_C, Nat.WithBot.one_le_iff_zero_lt] + +lemma eq_of_dvd_of_natDegree_le_of_leadingCoeff {p q : R[X]} (hpq : p ∣ q) + (h₁ : q.natDegree ≤ p.natDegree) (h₂ : p.leadingCoeff = q.leadingCoeff) : + p = q := by + by_cases hq : q = 0 + · rwa [hq, leadingCoeff_zero, leadingCoeff_eq_zero, ← hq] at h₂ + replace h₁ := (natDegree_le_of_dvd hpq hq).antisymm h₁ + obtain ⟨u, rfl⟩ := hpq + replace hq := mul_ne_zero_iff.mp hq + rw [natDegree_mul hq.1 hq.2, self_eq_add_right] at h₁ + rw [eq_C_of_natDegree_eq_zero h₁, leadingCoeff_mul, leadingCoeff_C, + eq_comm, mul_eq_left₀ (leadingCoeff_ne_zero.mpr hq.1)] at h₂ + rw [eq_C_of_natDegree_eq_zero h₁, h₂, map_one, mul_one] + +lemma associated_of_dvd_of_natDegree_le_of_leadingCoeff {p q : R[X]} (hpq : p ∣ q) + (h₁ : q.natDegree ≤ p.natDegree) (h₂ : q.leadingCoeff ∣ p.leadingCoeff) : + Associated p q := + have ⟨r, hr⟩ := hpq + have ⟨u, hu⟩ := associated_of_dvd_dvd ⟨leadingCoeff r, hr ▸ leadingCoeff_mul p r⟩ h₂ + ⟨Units.map C.toMonoidHom u, eq_of_dvd_of_natDegree_le_of_leadingCoeff + (by rwa [Units.mul_right_dvd]) (by simpa [natDegree_mul_C] using h₁) (by simpa using hu)⟩ + +lemma associated_of_dvd_of_natDegree_le {K} [Field K] {p q : K[X]} (hpq : p ∣ q) (hq : q ≠ 0) + (h₁ : q.natDegree ≤ p.natDegree) : Associated p q := + associated_of_dvd_of_natDegree_le_of_leadingCoeff hpq h₁ + (IsUnit.dvd (by rwa [← leadingCoeff_ne_zero, ← isUnit_iff_ne_zero] at hq)) + +lemma associated_of_dvd_of_degree_eq {K} [Field K] {p q : K[X]} (hpq : p ∣ q) + (h₁ : p.degree = q.degree) : Associated p q := + (Classical.em (q = 0)).elim (fun hq ↦ (show p = q by simpa [hq] using h₁) ▸ Associated.refl p) + (associated_of_dvd_of_natDegree_le hpq · (natDegree_le_natDegree h₁.ge)) + +lemma eq_leadingCoeff_mul_of_monic_of_dvd_of_natDegree_le {R} [CommRing R] {p q : R[X]} + (hp : p.Monic) (hdiv : p ∣ q) (hdeg : q.natDegree ≤ p.natDegree) : + q = C q.leadingCoeff * p := by + obtain ⟨r, hr⟩ := hdiv + obtain rfl | hq := eq_or_ne q 0; · simp + have rzero : r ≠ 0 := fun h => by simp [h, hq] at hr + rw [hr, natDegree_mul'] at hdeg; swap + · rw [hp.leadingCoeff, one_mul, leadingCoeff_ne_zero] + exact rzero + rw [mul_comm, @eq_C_of_natDegree_eq_zero _ _ r] at hr + · convert hr + convert leadingCoeff_C (coeff r 0) using 1 + rw [hr, leadingCoeff_mul_monic hp] + · exact (add_right_inj _).1 (le_antisymm hdeg <| Nat.le.intro rfl) + +lemma eq_of_monic_of_dvd_of_natDegree_le {R} [CommRing R] {p q : R[X]} (hp : p.Monic) + (hq : q.Monic) (hdiv : p ∣ q) (hdeg : q.natDegree ≤ p.natDegree) : q = p := by + convert eq_leadingCoeff_mul_of_monic_of_dvd_of_natDegree_le hp hdiv hdeg + rw [hq.leadingCoeff, C_1, one_mul] + end CommRing end Polynomial diff --git a/Mathlib/Algebra/Polynomial/Eval.lean b/Mathlib/Algebra/Polynomial/Eval.lean index e6b4d34e657aa..3c5b199f6c28d 100644 --- a/Mathlib/Algebra/Polynomial/Eval.lean +++ b/Mathlib/Algebra/Polynomial/Eval.lean @@ -1038,6 +1038,14 @@ theorem eval_eq_zero_of_dvd_of_eval_eq_zero : p ∣ q → eval x p = 0 → eval theorem eval_geom_sum {R} [CommSemiring R] {n : ℕ} {x : R} : eval x (∑ i ∈ range n, X ^ i) = ∑ i ∈ range n, x ^ i := by simp [eval_finset_sum] +variable [NoZeroDivisors R] + +lemma root_mul : IsRoot (p * q) a ↔ IsRoot p a ∨ IsRoot q a := by + simp_rw [IsRoot, eval_mul, mul_eq_zero] + +lemma root_or_root_of_root_mul (h : IsRoot (p * q) a) : IsRoot p a ∨ IsRoot q a := + root_mul.1 h + end end Eval @@ -1155,4 +1163,50 @@ alias mul_X_sub_int_cast_comp := mul_X_sub_intCast_comp end Ring +section +variable [Semiring R] [CommRing S] [IsDomain S] (φ : R →+* S) {f : R[X]} + +lemma isUnit_of_isUnit_leadingCoeff_of_isUnit_map (hf : IsUnit f.leadingCoeff) + (H : IsUnit (map φ f)) : IsUnit f := by + have dz := degree_eq_zero_of_isUnit H + rw [degree_map_eq_of_leadingCoeff_ne_zero] at dz + · rw [eq_C_of_degree_eq_zero dz] + refine IsUnit.map C ?_ + convert hf + change coeff f 0 = coeff f (natDegree f) + rw [(degree_eq_iff_natDegree_eq _).1 dz] + · rfl + rintro rfl + simp at H + · intro h + have u : IsUnit (φ f.leadingCoeff) := IsUnit.map φ hf + rw [h] at u + simp at u + +end + +section +variable [CommRing R] [IsDomain R] [CommRing S] [IsDomain S] (φ : R →+* S) + +/-- A polynomial over an integral domain `R` is irreducible if it is monic and +irreducible after mapping into an integral domain `S`. + +A special case of this lemma is that a polynomial over `ℤ` is irreducible if +it is monic and irreducible over `ℤ/pℤ` for some prime `p`. +-/ +lemma Monic.irreducible_of_irreducible_map (f : R[X]) (h_mon : Monic f) + (h_irr : Irreducible (f.map φ)) : Irreducible f := by + refine ⟨h_irr.not_unit ∘ IsUnit.map (mapRingHom φ), fun a b h => ?_⟩ + dsimp [Monic] at h_mon + have q := (leadingCoeff_mul a b).symm + rw [← h, h_mon] at q + refine (h_irr.isUnit_or_isUnit <| + (congr_arg (Polynomial.map φ) h).trans (Polynomial.map_mul φ)).imp ?_ ?_ <;> + apply isUnit_of_isUnit_leadingCoeff_of_isUnit_map <;> + apply isUnit_of_mul_eq_one + · exact q + · rw [mul_comm] + exact q + +end end Polynomial diff --git a/Mathlib/Algebra/Polynomial/Mirror.lean b/Mathlib/Algebra/Polynomial/Mirror.lean index def8942a33b30..9f3b754cd0a0a 100644 --- a/Mathlib/Algebra/Polynomial/Mirror.lean +++ b/Mathlib/Algebra/Polynomial/Mirror.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Thomas Browning -/ import Mathlib.Algebra.BigOperators.NatAntidiagonal -import Mathlib.Algebra.Polynomial.RingDivision +import Mathlib.Algebra.Polynomial.Reverse /-! # "Mirror" of a univariate polynomial diff --git a/Mathlib/Algebra/Polynomial/Monic.lean b/Mathlib/Algebra/Polynomial/Monic.lean index eb0f0cf67b6b0..9b7ed2ba528e5 100644 --- a/Mathlib/Algebra/Polynomial/Monic.lean +++ b/Mathlib/Algebra/Polynomial/Monic.lean @@ -130,6 +130,19 @@ theorem Monic.of_mul_monic_right (hq : q.Monic) (hpq : (p * q).Monic) : p.Monic namespace Monic +lemma comp (hp : p.Monic) (hq : q.Monic) (h : q.natDegree ≠ 0) : (p.comp q).Monic := by + nontriviality R + have : (p.comp q).natDegree = p.natDegree * q.natDegree := + natDegree_comp_eq_of_mul_ne_zero <| by simp [hp.leadingCoeff, hq.leadingCoeff] + rw [Monic.def, Polynomial.leadingCoeff, this, coeff_comp_degree_mul_degree h, hp.leadingCoeff, + hq.leadingCoeff, one_pow, mul_one] + +lemma comp_X_add_C (hp : p.Monic) (r : R) : (p.comp (X + C r)).Monic := by + nontriviality R + refine hp.comp (monic_X_add_C _) fun ha ↦ ?_ + rw [natDegree_X_add_C] at ha + exact one_ne_zero ha + @[simp] theorem natDegree_eq_zero_iff_eq_one (hp : p.Monic) : p.natDegree = 0 ↔ p = 1 := by constructor <;> intro h @@ -262,6 +275,78 @@ theorem Monic.nextCoeff_prod (s : Finset ι) (f : ι → R[X]) (h : ∀ i ∈ s, nextCoeff (∏ i ∈ s, f i) = ∑ i ∈ s, nextCoeff (f i) := Monic.nextCoeff_multiset_prod s.1 f h +variable [NoZeroDivisors R] {p q : R[X]} + +lemma irreducible_of_monic (hp : p.Monic) (hp1 : p ≠ 1) : + Irreducible p ↔ ∀ f g : R[X], f.Monic → g.Monic → f * g = p → f = 1 ∨ g = 1 := by + refine + ⟨fun h f g hf hg hp => (h.2 f g hp.symm).imp hf.eq_one_of_isUnit hg.eq_one_of_isUnit, fun h => + ⟨hp1 ∘ hp.eq_one_of_isUnit, fun f g hfg => + (h (g * C f.leadingCoeff) (f * C g.leadingCoeff) ?_ ?_ ?_).symm.imp + (isUnit_of_mul_eq_one f _) + (isUnit_of_mul_eq_one g _)⟩⟩ + · rwa [Monic, leadingCoeff_mul, leadingCoeff_C, ← leadingCoeff_mul, mul_comm, ← hfg, ← Monic] + · rwa [Monic, leadingCoeff_mul, leadingCoeff_C, ← leadingCoeff_mul, ← hfg, ← Monic] + · rw [mul_mul_mul_comm, ← C_mul, ← leadingCoeff_mul, ← hfg, hp.leadingCoeff, C_1, mul_one, + mul_comm, ← hfg] + + +lemma Monic.irreducible_iff_natDegree (hp : p.Monic) : + Irreducible p ↔ + p ≠ 1 ∧ ∀ f g : R[X], f.Monic → g.Monic → f * g = p → f.natDegree = 0 ∨ g.natDegree = 0 := by + by_cases hp1 : p = 1; · simp [hp1] + rw [irreducible_of_monic hp hp1, and_iff_right hp1] + refine forall₄_congr fun a b ha hb => ?_ + rw [ha.natDegree_eq_zero_iff_eq_one, hb.natDegree_eq_zero_iff_eq_one] + +lemma Monic.irreducible_iff_natDegree' (hp : p.Monic) : Irreducible p ↔ p ≠ 1 ∧ + ∀ f g : R[X], f.Monic → g.Monic → f * g = p → g.natDegree ∉ Ioc 0 (p.natDegree / 2) := by + simp_rw [hp.irreducible_iff_natDegree, mem_Ioc, Nat.le_div_iff_mul_le zero_lt_two, mul_two] + apply and_congr_right' + constructor <;> intro h f g hf hg he <;> subst he + · rw [hf.natDegree_mul hg, add_le_add_iff_right] + exact fun ha => (h f g hf hg rfl).elim (ha.1.trans_le ha.2).ne' ha.1.ne' + · simp_rw [hf.natDegree_mul hg, pos_iff_ne_zero] at h + contrapose! h + obtain hl | hl := le_total f.natDegree g.natDegree + · exact ⟨g, f, hg, hf, mul_comm g f, h.1, add_le_add_left hl _⟩ + · exact ⟨f, g, hf, hg, rfl, h.2, add_le_add_right hl _⟩ + +/-- Alternate phrasing of `Polynomial.Monic.irreducible_iff_natDegree'` where we only have to check +one divisor at a time. -/ +lemma Monic.irreducible_iff_lt_natDegree_lt {p : R[X]} (hp : p.Monic) (hp1 : p ≠ 1) : + Irreducible p ↔ ∀ q, Monic q → natDegree q ∈ Finset.Ioc 0 (natDegree p / 2) → ¬ q ∣ p := by + rw [hp.irreducible_iff_natDegree', and_iff_right hp1] + constructor + · rintro h g hg hdg ⟨f, rfl⟩ + exact h f g (hg.of_mul_monic_left hp) hg (mul_comm f g) hdg + · rintro h f g - hg rfl hdg + exact h g hg hdg (dvd_mul_left g f) + +lemma Monic.not_irreducible_iff_exists_add_mul_eq_coeff (hm : p.Monic) (hnd : p.natDegree = 2) : + ¬Irreducible p ↔ ∃ c₁ c₂, p.coeff 0 = c₁ * c₂ ∧ p.coeff 1 = c₁ + c₂ := by + cases subsingleton_or_nontrivial R + · simp [natDegree_of_subsingleton] at hnd + rw [hm.irreducible_iff_natDegree', and_iff_right, hnd] + · push_neg + constructor + · rintro ⟨a, b, ha, hb, rfl, hdb⟩ + simp only [zero_lt_two, Nat.div_self, Nat.Ioc_succ_singleton, zero_add, mem_singleton] at hdb + have hda := hnd + rw [ha.natDegree_mul hb, hdb] at hda + use a.coeff 0, b.coeff 0, mul_coeff_zero a b + simpa only [nextCoeff, hnd, add_right_cancel hda, hdb] using ha.nextCoeff_mul hb + · rintro ⟨c₁, c₂, hmul, hadd⟩ + refine + ⟨X + C c₁, X + C c₂, monic_X_add_C _, monic_X_add_C _, ?_, ?_⟩ + · rw [p.as_sum_range_C_mul_X_pow, hnd, Finset.sum_range_succ, Finset.sum_range_succ, + Finset.sum_range_one, ← hnd, hm.coeff_natDegree, hnd, hmul, hadd, C_mul, C_add, C_1] + ring + · rw [mem_Ioc, natDegree_X_add_C _] + simp + · rintro rfl + simp [natDegree_one] at hnd + end CommSemiring section Semiring @@ -353,6 +438,9 @@ theorem not_isUnit_X_pow_sub_one (R : Type*) [CommRing R] [Nontrivial R] (n : apply hn rw [← @natDegree_one R, ← (monic_X_pow_sub_C _ hn).eq_one_of_isUnit h, natDegree_X_pow_sub_C] +lemma 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) + theorem Monic.sub_of_left {p q : R[X]} (hp : Monic p) (hpq : degree q < degree p) : Monic (p - q) := by rw [sub_eq_add_neg] diff --git a/Mathlib/Algebra/Polynomial/RingDivision.lean b/Mathlib/Algebra/Polynomial/RingDivision.lean index b93ceffe37ce4..c557bfff9de6c 100644 --- a/Mathlib/Algebra/Polynomial/RingDivision.lean +++ b/Mathlib/Algebra/Polynomial/RingDivision.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Johannes Hölzl, Kim Morrison, Jens Wagemaker, Johan Commelin -/ import Mathlib.Algebra.Polynomial.AlgebraMap -import Mathlib.Algebra.Polynomial.BigOperators import Mathlib.Algebra.Polynomial.Degree.Lemmas import Mathlib.Algebra.Polynomial.Div @@ -27,26 +26,6 @@ universe u v w z variable {R : Type u} {S : Type v} {T : Type w} {a b : R} {n : ℕ} -section Semiring - -variable [Semiring R] {p q : R[X]} - -theorem Monic.comp (hp : p.Monic) (hq : q.Monic) (h : q.natDegree ≠ 0) : (p.comp q).Monic := by - nontriviality R - have : (p.comp q).natDegree = p.natDegree * q.natDegree := by - apply natDegree_comp_eq_of_mul_ne_zero - simp [hp.leadingCoeff, hq.leadingCoeff] - rw [Monic.def, Polynomial.leadingCoeff, this, coeff_comp_degree_mul_degree h, hp.leadingCoeff, - hq.leadingCoeff, one_pow, mul_one] - -theorem Monic.comp_X_add_C (hp : p.Monic) (r : R) : (p.comp (X + C r)).Monic := by - nontriviality R - refine hp.comp (monic_X_add_C _) fun ha => ?_ - rw [natDegree_X_add_C] at ha - exact one_ne_zero ha - -end Semiring - section CommRing variable [CommRing R] {p q : R[X]} @@ -63,26 +42,6 @@ theorem degree_pos_of_aeval_root [Algebra R S] {p : R[X]} (hp : p ≠ 0) {z : S} (inj : ∀ x : R, algebraMap R S x = 0 → x = 0) : 0 < p.degree := natDegree_pos_iff_degree_pos.mp (natDegree_pos_of_aeval_root hp hz inj) -theorem modByMonic_eq_of_dvd_sub (hq : q.Monic) {p₁ p₂ : R[X]} (h : q ∣ p₁ - p₂) : - p₁ %ₘ q = p₂ %ₘ q := by - nontriviality R - obtain ⟨f, sub_eq⟩ := h - refine (div_modByMonic_unique (p₂ /ₘ q + f) _ hq ⟨?_, degree_modByMonic_lt _ hq⟩).2 - rw [sub_eq_iff_eq_add.mp sub_eq, mul_add, ← add_assoc, modByMonic_add_div _ hq, add_comm] - -theorem add_modByMonic (p₁ p₂ : R[X]) : (p₁ + p₂) %ₘ q = p₁ %ₘ q + p₂ %ₘ q := by - by_cases hq : q.Monic - · cases' subsingleton_or_nontrivial R with hR hR - · simp only [eq_iff_true_of_subsingleton] - · exact - (div_modByMonic_unique (p₁ /ₘ q + p₂ /ₘ q) _ hq - ⟨by - rw [mul_add, add_left_comm, add_assoc, modByMonic_add_div _ hq, ← add_assoc, - add_comm (q * _), modByMonic_add_div _ hq], - (degree_add_le _ _).trans_lt - (max_lt (degree_modByMonic_lt _ hq) (degree_modByMonic_lt _ hq))⟩).2 - · simp_rw [modByMonic_eq_of_not_monic _ hq] - theorem smul_modByMonic (c : R) (p : R[X]) : c • p %ₘ q = c • (p %ₘ q) := by by_cases hq : q.Monic · cases' subsingleton_or_nontrivial R with hR hR @@ -100,12 +59,6 @@ def modByMonicHom (q : R[X]) : R[X] →ₗ[R] R[X] where map_add' := add_modByMonic map_smul' := smul_modByMonic -theorem neg_modByMonic (p mod : R[X]) : (-p) %ₘ mod = - (p %ₘ mod) := - (modByMonicHom mod).map_neg p - -theorem sub_modByMonic (a b mod : R[X]) : (a - b) %ₘ mod = a %ₘ mod - b %ₘ mod := - (modByMonicHom mod).map_sub a b - end section @@ -126,16 +79,6 @@ section NoZeroDivisors variable [Semiring R] [NoZeroDivisors R] {p q : R[X]} -instance : NoZeroDivisors R[X] where - eq_zero_or_eq_zero_of_mul_eq_zero h := by - rw [← leadingCoeff_eq_zero, ← leadingCoeff_eq_zero] - refine eq_zero_or_eq_zero_of_mul_eq_zero ?_ - rw [← leadingCoeff_zero, ← leadingCoeff_mul, h] - -theorem natDegree_mul (hp : p ≠ 0) (hq : q ≠ 0) : (p*q).natDegree = p.natDegree + q.natDegree := by - rw [← Nat.cast_inj (R := WithBot ℕ), ← degree_eq_natDegree (mul_ne_zero hp hq), - Nat.cast_add, ← degree_eq_natDegree hp, ← degree_eq_natDegree hq, degree_mul] - theorem trailingDegree_mul : (p * q).trailingDegree = p.trailingDegree + q.trailingDegree := by by_cases hp : p = 0 · rw [hp, zero_mul, trailingDegree_zero, top_add] @@ -145,307 +88,13 @@ theorem trailingDegree_mul : (p * q).trailingDegree = p.trailingDegree + q.trail trailingDegree_eq_natTrailingDegree (mul_ne_zero hp hq), natTrailingDegree_mul hp hq] apply WithTop.coe_add -@[simp] -theorem natDegree_pow (p : R[X]) (n : ℕ) : natDegree (p ^ n) = n * natDegree p := by - classical - obtain rfl | hp := eq_or_ne p 0 - · obtain rfl | hn := eq_or_ne n 0 <;> simp [*] - exact natDegree_pow' <| by - rw [← leadingCoeff_pow, Ne, leadingCoeff_eq_zero]; exact pow_ne_zero _ hp - -theorem degree_le_mul_left (p : R[X]) (hq : q ≠ 0) : degree p ≤ degree (p * q) := by - classical - exact if hp : p = 0 then by simp only [hp, zero_mul, le_refl] - else by - rw [degree_mul, degree_eq_natDegree hp, degree_eq_natDegree hq] - exact WithBot.coe_le_coe.2 (Nat.le_add_right _ _) - -theorem natDegree_le_of_dvd {p q : R[X]} (h1 : p ∣ q) (h2 : q ≠ 0) : p.natDegree ≤ q.natDegree := by - rcases h1 with ⟨q, rfl⟩; rw [mul_ne_zero_iff] at h2 - rw [natDegree_mul h2.1 h2.2]; exact Nat.le_add_right _ _ - -theorem degree_le_of_dvd {p q : R[X]} (h1 : p ∣ q) (h2 : q ≠ 0) : degree p ≤ degree q := by - rcases h1 with ⟨q, rfl⟩; rw [mul_ne_zero_iff] at h2 - exact degree_le_mul_left p h2.2 - -theorem eq_zero_of_dvd_of_degree_lt {p q : R[X]} (h₁ : p ∣ q) (h₂ : degree q < degree p) : - q = 0 := by - by_contra hc - exact (lt_iff_not_ge _ _).mp h₂ (degree_le_of_dvd h₁ hc) - -theorem eq_zero_of_dvd_of_natDegree_lt {p q : R[X]} (h₁ : p ∣ q) (h₂ : natDegree q < natDegree p) : - q = 0 := by - by_contra hc - exact (lt_iff_not_ge _ _).mp h₂ (natDegree_le_of_dvd h₁ hc) - -theorem not_dvd_of_degree_lt {p q : R[X]} (h0 : q ≠ 0) (hl : q.degree < p.degree) : ¬p ∣ q := by - by_contra hcontra - exact h0 (eq_zero_of_dvd_of_degree_lt hcontra hl) - -theorem not_dvd_of_natDegree_lt {p q : R[X]} (h0 : q ≠ 0) (hl : q.natDegree < p.natDegree) : - ¬p ∣ q := by - by_contra hcontra - exact h0 (eq_zero_of_dvd_of_natDegree_lt hcontra hl) - -/-- This lemma is useful for working with the `intDegree` of a rational function. -/ -theorem natDegree_sub_eq_of_prod_eq {p₁ p₂ q₁ q₂ : R[X]} (hp₁ : p₁ ≠ 0) (hq₁ : q₁ ≠ 0) - (hp₂ : p₂ ≠ 0) (hq₂ : q₂ ≠ 0) (h_eq : p₁ * q₂ = p₂ * q₁) : - (p₁.natDegree : ℤ) - q₁.natDegree = (p₂.natDegree : ℤ) - q₂.natDegree := by - rw [sub_eq_sub_iff_add_eq_add] - norm_cast - rw [← natDegree_mul hp₁ hq₂, ← natDegree_mul hp₂ hq₁, h_eq] - -theorem natDegree_eq_zero_of_isUnit (h : IsUnit p) : natDegree p = 0 := by - nontriviality R - obtain ⟨q, hq⟩ := h.exists_right_inv - have := natDegree_mul (left_ne_zero_of_mul_eq_one hq) (right_ne_zero_of_mul_eq_one hq) - rw [hq, natDegree_one, eq_comm, add_eq_zero] at this - exact this.1 - -theorem degree_eq_zero_of_isUnit [Nontrivial R] (h : IsUnit p) : degree p = 0 := - (natDegree_eq_zero_iff_degree_le_zero.mp <| natDegree_eq_zero_of_isUnit h).antisymm - (zero_le_degree_iff.mpr h.ne_zero) - -@[simp] -theorem degree_coe_units [Nontrivial R] (u : R[X]ˣ) : degree (u : R[X]) = 0 := - degree_eq_zero_of_isUnit ⟨u, rfl⟩ - -/-- Characterization of a unit of a polynomial ring over an integral domain `R`. -See `Polynomial.isUnit_iff_coeff_isUnit_isNilpotent` when `R` is a commutative ring. -/ -theorem isUnit_iff : IsUnit p ↔ ∃ r : R, IsUnit r ∧ C r = p := - ⟨fun hp => - ⟨p.coeff 0, - let h := eq_C_of_natDegree_eq_zero (natDegree_eq_zero_of_isUnit hp) - ⟨isUnit_C.1 (h ▸ hp), h.symm⟩⟩, - fun ⟨_, hr, hrp⟩ => hrp ▸ isUnit_C.2 hr⟩ - -theorem not_isUnit_of_degree_pos (p : R[X]) - (hpl : 0 < p.degree) : ¬ IsUnit p := by - cases subsingleton_or_nontrivial R - · simp [Subsingleton.elim p 0] at hpl - intro h - simp [degree_eq_zero_of_isUnit h] at hpl - -theorem not_isUnit_of_natDegree_pos (p : R[X]) - (hpl : 0 < p.natDegree) : ¬ IsUnit p := - not_isUnit_of_degree_pos _ (natDegree_pos_iff_degree_pos.mp hpl) - -end NoZeroDivisors - -section NoZeroDivisors - -variable [CommSemiring R] [NoZeroDivisors R] {p q : R[X]} - -theorem irreducible_of_monic (hp : p.Monic) (hp1 : p ≠ 1) : - Irreducible p ↔ ∀ f g : R[X], f.Monic → g.Monic → f * g = p → f = 1 ∨ g = 1 := by - refine - ⟨fun h f g hf hg hp => (h.2 f g hp.symm).imp hf.eq_one_of_isUnit hg.eq_one_of_isUnit, fun h => - ⟨hp1 ∘ hp.eq_one_of_isUnit, fun f g hfg => - (h (g * C f.leadingCoeff) (f * C g.leadingCoeff) ?_ ?_ ?_).symm.imp - (isUnit_of_mul_eq_one f _) - (isUnit_of_mul_eq_one g _)⟩⟩ - · rwa [Monic, leadingCoeff_mul, leadingCoeff_C, ← leadingCoeff_mul, mul_comm, ← hfg, ← Monic] - · rwa [Monic, leadingCoeff_mul, leadingCoeff_C, ← leadingCoeff_mul, ← hfg, ← Monic] - · rw [mul_mul_mul_comm, ← C_mul, ← leadingCoeff_mul, ← hfg, hp.leadingCoeff, C_1, mul_one, - mul_comm, ← hfg] - -theorem Monic.irreducible_iff_natDegree (hp : p.Monic) : - Irreducible p ↔ - p ≠ 1 ∧ ∀ f g : R[X], f.Monic → g.Monic → f * g = p → f.natDegree = 0 ∨ g.natDegree = 0 := by - by_cases hp1 : p = 1; · simp [hp1] - rw [irreducible_of_monic hp hp1, and_iff_right hp1] - refine forall₄_congr fun a b ha hb => ?_ - rw [ha.natDegree_eq_zero_iff_eq_one, hb.natDegree_eq_zero_iff_eq_one] - -theorem Monic.irreducible_iff_natDegree' (hp : p.Monic) : Irreducible p ↔ p ≠ 1 ∧ - ∀ f g : R[X], f.Monic → g.Monic → f * g = p → g.natDegree ∉ Ioc 0 (p.natDegree / 2) := by - simp_rw [hp.irreducible_iff_natDegree, mem_Ioc, Nat.le_div_iff_mul_le zero_lt_two, mul_two] - apply and_congr_right' - constructor <;> intro h f g hf hg he <;> subst he - · rw [hf.natDegree_mul hg, add_le_add_iff_right] - exact fun ha => (h f g hf hg rfl).elim (ha.1.trans_le ha.2).ne' ha.1.ne' - · simp_rw [hf.natDegree_mul hg, pos_iff_ne_zero] at h - contrapose! h - obtain hl | hl := le_total f.natDegree g.natDegree - · exact ⟨g, f, hg, hf, mul_comm g f, h.1, add_le_add_left hl _⟩ - · exact ⟨f, g, hf, hg, rfl, h.2, add_le_add_right hl _⟩ - -/-- Alternate phrasing of `Polynomial.Monic.irreducible_iff_natDegree'` where we only have to check -one divisor at a time. -/ -theorem Monic.irreducible_iff_lt_natDegree_lt {p : R[X]} (hp : p.Monic) (hp1 : p ≠ 1) : - Irreducible p ↔ ∀ q, Monic q → natDegree q ∈ Finset.Ioc 0 (natDegree p / 2) → ¬ q ∣ p := by - rw [hp.irreducible_iff_natDegree', and_iff_right hp1] - constructor - · rintro h g hg hdg ⟨f, rfl⟩ - exact h f g (hg.of_mul_monic_left hp) hg (mul_comm f g) hdg - · rintro h f g - hg rfl hdg - exact h g hg hdg (dvd_mul_left g f) - -theorem Monic.not_irreducible_iff_exists_add_mul_eq_coeff (hm : p.Monic) (hnd : p.natDegree = 2) : - ¬Irreducible p ↔ ∃ c₁ c₂, p.coeff 0 = c₁ * c₂ ∧ p.coeff 1 = c₁ + c₂ := by - cases subsingleton_or_nontrivial R - · simp [natDegree_of_subsingleton] at hnd - rw [hm.irreducible_iff_natDegree', and_iff_right, hnd] - · push_neg - constructor - · rintro ⟨a, b, ha, hb, rfl, hdb⟩ - simp only [zero_lt_two, Nat.div_self, Nat.Ioc_succ_singleton, zero_add, mem_singleton] at hdb - have hda := hnd - rw [ha.natDegree_mul hb, hdb] at hda - use a.coeff 0, b.coeff 0, mul_coeff_zero a b - simpa only [nextCoeff, hnd, add_right_cancel hda, hdb] using ha.nextCoeff_mul hb - · rintro ⟨c₁, c₂, hmul, hadd⟩ - refine - ⟨X + C c₁, X + C c₂, monic_X_add_C _, monic_X_add_C _, ?_, ?_⟩ - · rw [p.as_sum_range_C_mul_X_pow, hnd, Finset.sum_range_succ, Finset.sum_range_succ, - Finset.sum_range_one, ← hnd, hm.coeff_natDegree, hnd, hmul, hadd, C_mul, C_add, C_1] - ring - · rw [mem_Ioc, natDegree_X_add_C _] - simp - · rintro rfl - simp [natDegree_one] at hnd - -theorem root_mul : IsRoot (p * q) a ↔ IsRoot p a ∨ IsRoot q a := by - simp_rw [IsRoot, eval_mul, mul_eq_zero] - -theorem root_or_root_of_root_mul (h : IsRoot (p * q) a) : IsRoot p a ∨ IsRoot q a := - root_mul.1 h - end NoZeroDivisors -section Ring - -variable [Ring R] [IsDomain R] {p q : R[X]} - -instance : IsDomain R[X] := - NoZeroDivisors.to_isDomain _ - -end Ring - -section CommSemiring - -variable [CommSemiring R] {a p : R[X]} - -section Monic - -variable (hp : p.Monic) -include hp - -theorem Monic.C_dvd_iff_isUnit {a : R} : C a ∣ p ↔ IsUnit a := - ⟨fun h => isUnit_iff_dvd_one.mpr <| - hp.coeff_natDegree ▸ (C_dvd_iff_dvd_coeff _ _).mp h p.natDegree, - fun ha => (ha.map C).dvd⟩ - -theorem Monic.natDegree_pos : 0 < natDegree p ↔ p ≠ 1 := - Nat.pos_iff_ne_zero.trans hp.natDegree_eq_zero.not - -theorem Monic.degree_pos : 0 < degree p ↔ p ≠ 1 := - natDegree_pos_iff_degree_pos.symm.trans hp.natDegree_pos - -theorem Monic.degree_pos_of_not_isUnit (hu : ¬IsUnit p) : 0 < degree p := - hp.degree_pos.mpr (fun hp' ↦ (hp' ▸ hu) isUnit_one) - -theorem Monic.natDegree_pos_of_not_isUnit (hu : ¬IsUnit p) : 0 < natDegree p := - hp.natDegree_pos.mpr (fun hp' ↦ (hp' ▸ hu) isUnit_one) - -theorem degree_pos_of_not_isUnit_of_dvd_monic (ha : ¬IsUnit a) (hap : a ∣ p) : 0 < degree a := by - contrapose! ha with h - rw [Polynomial.eq_C_of_degree_le_zero h] at hap ⊢ - simpa [hp.C_dvd_iff_isUnit, isUnit_C] using hap - -theorem natDegree_pos_of_not_isUnit_of_dvd_monic (ha : ¬IsUnit a) (hap : a ∣ p) : 0 < natDegree a := - natDegree_pos_iff_degree_pos.mpr <| degree_pos_of_not_isUnit_of_dvd_monic hp ha hap - -end Monic - -theorem eq_zero_of_mul_eq_zero_of_smul (P : R[X]) (h : ∀ r : R, r • P = 0 → r = 0) : - ∀ (Q : R[X]), P * Q = 0 → Q = 0 := by - intro Q hQ - suffices ∀ i, P.coeff i • Q = 0 by - rw [← leadingCoeff_eq_zero] - apply h - simpa [ext_iff, mul_comm Q.leadingCoeff] using fun i ↦ congr_arg (·.coeff Q.natDegree) (this i) - apply Nat.strong_decreasing_induction - · use P.natDegree - intro i hi - rw [coeff_eq_zero_of_natDegree_lt hi, zero_smul] - intro l IH - obtain _|hl := (natDegree_smul_le (P.coeff l) Q).lt_or_eq - · apply eq_zero_of_mul_eq_zero_of_smul _ h (P.coeff l • Q) - rw [smul_eq_C_mul, mul_left_comm, hQ, mul_zero] - suffices P.coeff l * Q.leadingCoeff = 0 by - rwa [← leadingCoeff_eq_zero, ← coeff_natDegree, coeff_smul, hl, coeff_natDegree, smul_eq_mul] - let m := Q.natDegree - suffices (P * Q).coeff (l + m) = P.coeff l * Q.leadingCoeff by rw [← this, hQ, coeff_zero] - rw [coeff_mul] - apply Finset.sum_eq_single (l, m) _ (by simp) - simp only [Finset.mem_antidiagonal, ne_eq, Prod.forall, Prod.mk.injEq, not_and] - intro i j hij H - obtain hi|rfl|hi := lt_trichotomy i l - · have hj : m < j := by omega - rw [coeff_eq_zero_of_natDegree_lt hj, mul_zero] - · omega - · rw [← coeff_C_mul, ← smul_eq_C_mul, IH _ hi, coeff_zero] -termination_by Q => Q.natDegree - -open nonZeroDivisors in -/-- *McCoy theorem*: a polynomial `P : R[X]` is a zerodivisor if and only if there is `a : R` -such that `a ≠ 0` and `a • P = 0`. -/ -theorem nmem_nonZeroDivisors_iff {P : R[X]} : P ∉ R[X]⁰ ↔ ∃ a : R, a ≠ 0 ∧ a • P = 0 := by - refine ⟨fun hP ↦ ?_, fun ⟨a, ha, h⟩ h1 ↦ ha <| C_eq_zero.1 <| (h1 _) <| smul_eq_C_mul a ▸ h⟩ - by_contra! h - obtain ⟨Q, hQ⟩ := _root_.nmem_nonZeroDivisors_iff.1 hP - refine hQ.2 (eq_zero_of_mul_eq_zero_of_smul P (fun a ha ↦ ?_) Q (mul_comm P _ ▸ hQ.1)) - contrapose! ha - exact h a ha - -open nonZeroDivisors in -protected lemma mem_nonZeroDivisors_iff {P : R[X]} : P ∈ R[X]⁰ ↔ ∀ a : R, a • P = 0 → a = 0 := by - simpa [not_imp_not] using (nmem_nonZeroDivisors_iff (P := P)).not - -end CommSemiring section CommRing variable [CommRing R] -/- Porting note: the ML3 proof no longer worked because of a conflict in the -inferred type and synthesized type for `DecidableRel` when using `Nat.le_find_iff` from -`Mathlib.Algebra.Polynomial.Div` After some discussion on [Zulip] -(https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/decidability.20leakage) -introduced `Polynomial.rootMultiplicity_eq_nat_find_of_nonzero` to contain the issue --/ -/-- The multiplicity of `a` as root of a nonzero polynomial `p` is at least `n` iff - `(X - a) ^ n` divides `p`. -/ -theorem le_rootMultiplicity_iff {p : R[X]} (p0 : p ≠ 0) {a : R} {n : ℕ} : - n ≤ rootMultiplicity a p ↔ (X - C a) ^ n ∣ p := by - classical - rw [rootMultiplicity_eq_nat_find_of_nonzero p0, @Nat.le_find_iff _ (_)] - simp_rw [Classical.not_not] - refine ⟨fun h => ?_, fun h m hm => (pow_dvd_pow _ hm).trans h⟩ - cases' n with n - · rw [pow_zero] - apply one_dvd - · exact h n n.lt_succ_self - -theorem rootMultiplicity_le_iff {p : R[X]} (p0 : p ≠ 0) (a : R) (n : ℕ) : - rootMultiplicity a p ≤ n ↔ ¬(X - C a) ^ (n + 1) ∣ p := by - rw [← (le_rootMultiplicity_iff p0).not, not_le, Nat.lt_add_one_iff] - -theorem pow_rootMultiplicity_not_dvd {p : R[X]} (p0 : p ≠ 0) (a : R) : - ¬(X - C a) ^ (rootMultiplicity a p + 1) ∣ p := by rw [← rootMultiplicity_le_iff p0] - -theorem X_sub_C_pow_dvd_iff {p : R[X]} {t : R} {n : ℕ} : - (X - C t) ^ n ∣ p ↔ X ^ n ∣ p.comp (X + C t) := by - convert (map_dvd_iff <| algEquivAevalXAddC t).symm using 2 - simp [C_eq_algebraMap] - -theorem comp_X_add_C_eq_zero_iff {p : R[X]} (t : R) : - p.comp (X + C t) = 0 ↔ p = 0 := AddEquivClass.map_eq_zero_iff (algEquivAevalXAddC t) - -theorem comp_X_add_C_ne_zero_iff {p : R[X]} (t : R) : - p.comp (X + C t) ≠ 0 ↔ p ≠ 0 := Iff.not <| comp_X_add_C_eq_zero_iff t - theorem rootMultiplicity_eq_rootMultiplicity {p : R[X]} {t : R} : p.rootMultiplicity t = (p.comp (X + C t)).rootMultiplicity 0 := by classical @@ -455,35 +104,11 @@ theorem rootMultiplicity_eq_rootMultiplicity {p : R[X]} {t : R} : convert (multiplicity_map_eq <| algEquivAevalXAddC t).symm using 2 simp [C_eq_algebraMap] -theorem rootMultiplicity_eq_natTrailingDegree' {p : R[X]} : - p.rootMultiplicity 0 = p.natTrailingDegree := by - by_cases h : p = 0 - · simp only [h, rootMultiplicity_zero, natTrailingDegree_zero] - refine le_antisymm ?_ ?_ - · rw [rootMultiplicity_le_iff h, map_zero, sub_zero, X_pow_dvd_iff, not_forall] - exact ⟨p.natTrailingDegree, - fun h' ↦ trailingCoeff_nonzero_iff_nonzero.2 h <| h' <| Nat.lt.base _⟩ - · rw [le_rootMultiplicity_iff h, map_zero, sub_zero, X_pow_dvd_iff] - exact fun _ ↦ coeff_eq_zero_of_lt_natTrailingDegree - +/-- See `Polynomial.rootMultiplicity_eq_natTrailingDegree'` for the special case of `t = 0`. -/ theorem rootMultiplicity_eq_natTrailingDegree {p : R[X]} {t : R} : p.rootMultiplicity t = (p.comp (X + C t)).natTrailingDegree := rootMultiplicity_eq_rootMultiplicity.trans rootMultiplicity_eq_natTrailingDegree' -theorem eval_divByMonic_eq_trailingCoeff_comp {p : R[X]} {t : R} : - (p /ₘ (X - C t) ^ p.rootMultiplicity t).eval t = (p.comp (X + C t)).trailingCoeff := by - obtain rfl | hp := eq_or_ne p 0 - · rw [zero_divByMonic, eval_zero, zero_comp, trailingCoeff_zero] - have mul_eq := p.pow_mul_divByMonic_rootMultiplicity_eq t - set m := p.rootMultiplicity t - set g := p /ₘ (X - C t) ^ m - have : (g.comp (X + C t)).coeff 0 = g.eval t := by - rw [coeff_zero_eq_eval_zero, eval_comp, eval_add, eval_X, eval_C, zero_add] - rw [← congr_arg (comp · <| X + C t) mul_eq, mul_comp, pow_comp, sub_comp, X_comp, C_comp, - add_sub_cancel_right, ← reverse_leadingCoeff, reverse_X_pow_mul, reverse_leadingCoeff, - trailingCoeff, Nat.le_zero.1 (natTrailingDegree_le_of_ne_zero <| - this ▸ eval_divByMonic_pow_rootMultiplicity_ne_zero t hp), this] - section nonZeroDivisors open scoped nonZeroDivisors @@ -527,17 +152,6 @@ theorem rootMultiplicity_X_sub_C [Nontrivial R] [DecidableEq R] {x y : R} : exact rootMultiplicity_X_sub_C_self exact rootMultiplicity_eq_zero (mt root_X_sub_C.mp (Ne.symm hxy)) -/-- The multiplicity of `p + q` is at least the minimum of the multiplicities. -/ -theorem rootMultiplicity_add {p q : R[X]} (a : R) (hzero : p + q ≠ 0) : - min (rootMultiplicity a p) (rootMultiplicity a q) ≤ rootMultiplicity a (p + q) := by - rw [le_rootMultiplicity_iff hzero] - exact min_pow_dvd_add (pow_rootMultiplicity_dvd p a) (pow_rootMultiplicity_dvd q a) - -theorem le_rootMultiplicity_mul {p q : R[X]} (x : R) (hpq : p * q ≠ 0) : - rootMultiplicity x p + rootMultiplicity x q ≤ rootMultiplicity x (p * q) := by - rw [le_rootMultiplicity_iff hpq, pow_add] - exact mul_dvd_mul (pow_rootMultiplicity_dvd p x) (pow_rootMultiplicity_dvd q x) - theorem rootMultiplicity_mul' {p q : R[X]} {x : R} (hpq : (p /ₘ (X - C x) ^ p.rootMultiplicity x).eval x * (q /ₘ (X - C x) ^ q.rootMultiplicity x).eval x ≠ 0) : @@ -545,18 +159,6 @@ theorem rootMultiplicity_mul' {p q : R[X]} {x : R} simp_rw [eval_divByMonic_eq_trailingCoeff_comp] at hpq simp_rw [rootMultiplicity_eq_natTrailingDegree, mul_comp, natTrailingDegree_mul' hpq] -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] @@ -570,101 +172,10 @@ theorem Monic.neg_one_pow_natDegree_mul_comp_neg_X {p : R[X]} (hp : p.Monic) : variable [IsDomain R] {p q : R[X]} -@[simp] -theorem natDegree_coe_units (u : R[X]ˣ) : natDegree (u : R[X]) = 0 := - natDegree_eq_of_degree_eq_some (degree_coe_units u) - -theorem coeff_coe_units_zero_ne_zero (u : R[X]ˣ) : coeff (u : R[X]) 0 ≠ 0 := by - conv in 0 => rw [← natDegree_coe_units u] - rw [← leadingCoeff, Ne, leadingCoeff_eq_zero] - exact Units.ne_zero _ - theorem degree_eq_degree_of_associated (h : Associated p q) : degree p = degree q := by let ⟨u, hu⟩ := h simp [hu.symm] -theorem degree_eq_one_of_irreducible_of_root (hi : Irreducible p) {x : R} (hx : IsRoot p x) : - degree p = 1 := - let ⟨g, hg⟩ := dvd_iff_isRoot.2 hx - have : IsUnit (X - C x) ∨ IsUnit g := hi.isUnit_or_isUnit hg - this.elim - (fun h => by - have h₁ : degree (X - C x) = 1 := degree_X_sub_C x - have h₂ : degree (X - C x) = 0 := degree_eq_zero_of_isUnit h - rw [h₁] at h₂; exact absurd h₂ (by decide)) - fun hgu => by rw [hg, degree_mul, degree_X_sub_C, degree_eq_zero_of_isUnit hgu, add_zero] - -/-- Division by a monic polynomial doesn't change the leading coefficient. -/ -theorem leadingCoeff_divByMonic_of_monic {R : Type u} [CommRing R] {p q : R[X]} (hmonic : q.Monic) - (hdegree : q.degree ≤ p.degree) : (p /ₘ q).leadingCoeff = p.leadingCoeff := by - nontriviality - have h : q.leadingCoeff * (p /ₘ q).leadingCoeff ≠ 0 := by - simpa [divByMonic_eq_zero_iff hmonic, hmonic.leadingCoeff, - Nat.WithBot.one_le_iff_zero_lt] using hdegree - nth_rw 2 [← modByMonic_add_div p hmonic] - rw [leadingCoeff_add_of_degree_lt, leadingCoeff_monic_mul hmonic] - rw [degree_mul' h, degree_add_divByMonic hmonic hdegree] - exact (degree_modByMonic_lt p hmonic).trans_le hdegree - -theorem leadingCoeff_divByMonic_X_sub_C (p : R[X]) (hp : degree p ≠ 0) (a : R) : - leadingCoeff (p /ₘ (X - C a)) = leadingCoeff p := by - nontriviality - cases' hp.lt_or_lt with hd hd - · rw [degree_eq_bot.mp <| Nat.WithBot.lt_zero_iff.mp hd, zero_divByMonic] - refine leadingCoeff_divByMonic_of_monic (monic_X_sub_C a) ?_ - rwa [degree_X_sub_C, Nat.WithBot.one_le_iff_zero_lt] - -theorem eq_of_dvd_of_natDegree_le_of_leadingCoeff {p q : R[X]} (hpq : p ∣ q) - (h₁ : q.natDegree ≤ p.natDegree) (h₂ : p.leadingCoeff = q.leadingCoeff) : - p = q := by - by_cases hq : q = 0 - · rwa [hq, leadingCoeff_zero, leadingCoeff_eq_zero, ← hq] at h₂ - replace h₁ := (natDegree_le_of_dvd hpq hq).antisymm h₁ - obtain ⟨u, rfl⟩ := hpq - replace hq := mul_ne_zero_iff.mp hq - rw [natDegree_mul hq.1 hq.2, self_eq_add_right] at h₁ - rw [eq_C_of_natDegree_eq_zero h₁, leadingCoeff_mul, leadingCoeff_C, - eq_comm, mul_eq_left₀ (leadingCoeff_ne_zero.mpr hq.1)] at h₂ - rw [eq_C_of_natDegree_eq_zero h₁, h₂, map_one, mul_one] - -theorem associated_of_dvd_of_natDegree_le_of_leadingCoeff {p q : R[X]} (hpq : p ∣ q) - (h₁ : q.natDegree ≤ p.natDegree) (h₂ : q.leadingCoeff ∣ p.leadingCoeff) : - Associated p q := - have ⟨r, hr⟩ := hpq - have ⟨u, hu⟩ := associated_of_dvd_dvd ⟨leadingCoeff r, hr ▸ leadingCoeff_mul p r⟩ h₂ - ⟨Units.map C.toMonoidHom u, eq_of_dvd_of_natDegree_le_of_leadingCoeff - (by rwa [Units.mul_right_dvd]) (by simpa [natDegree_mul_C] using h₁) (by simpa using hu)⟩ - -theorem associated_of_dvd_of_natDegree_le {K} [Field K] {p q : K[X]} (hpq : p ∣ q) (hq : q ≠ 0) - (h₁ : q.natDegree ≤ p.natDegree) : Associated p q := - associated_of_dvd_of_natDegree_le_of_leadingCoeff hpq h₁ - (IsUnit.dvd (by rwa [← leadingCoeff_ne_zero, ← isUnit_iff_ne_zero] at hq)) - -theorem associated_of_dvd_of_degree_eq {K} [Field K] {p q : K[X]} (hpq : p ∣ q) - (h₁ : p.degree = q.degree) : Associated p q := - (Classical.em (q = 0)).elim (fun hq ↦ (show p = q by simpa [hq] using h₁) ▸ Associated.refl p) - (associated_of_dvd_of_natDegree_le hpq · (natDegree_le_natDegree h₁.ge)) - -theorem eq_leadingCoeff_mul_of_monic_of_dvd_of_natDegree_le {R} [CommRing R] {p q : R[X]} - (hp : p.Monic) (hdiv : p ∣ q) (hdeg : q.natDegree ≤ p.natDegree) : - q = C q.leadingCoeff * p := by - obtain ⟨r, hr⟩ := hdiv - obtain rfl | hq := eq_or_ne q 0; · simp - have rzero : r ≠ 0 := fun h => by simp [h, hq] at hr - rw [hr, natDegree_mul'] at hdeg; swap - · rw [hp.leadingCoeff, one_mul, leadingCoeff_ne_zero] - exact rzero - rw [mul_comm, @eq_C_of_natDegree_eq_zero _ _ r] at hr - · convert hr - convert leadingCoeff_C (coeff r 0) using 1 - rw [hr, leadingCoeff_mul_monic hp] - · exact (add_right_inj _).1 (le_antisymm hdeg <| Nat.le.intro rfl) - -theorem eq_of_monic_of_dvd_of_natDegree_le {R} [CommRing R] {p q : R[X]} (hp : p.Monic) - (hq : q.Monic) (hdiv : p ∣ q) (hdeg : q.natDegree ≤ p.natDegree) : q = p := by - convert eq_leadingCoeff_mul_of_monic_of_dvd_of_natDegree_le hp hdiv hdeg - rw [hq.leadingCoeff, C_1, one_mul] - theorem prime_X_sub_C (r : R) : Prime (X - C r) := ⟨X_sub_C_ne_zero r, not_isUnit_X_sub_C r, fun _ _ => by simp_rw [dvd_iff_isRoot, IsRoot.def, eval_mul, mul_eq_zero] @@ -687,31 +198,6 @@ theorem irreducible_X : Irreducible (X : R[X]) := theorem Monic.irreducible_of_degree_eq_one (hp1 : degree p = 1) (hm : Monic p) : Irreducible p := (hm.prime_of_degree_eq_one hp1).irreducible -@[simp] -theorem natDegree_multiset_prod_X_sub_C_eq_card (s : Multiset R) : - (s.map fun a => X - C a).prod.natDegree = Multiset.card s := by - rw [natDegree_multiset_prod_of_monic, Multiset.map_map] - · simp only [(· ∘ ·), natDegree_X_sub_C, Multiset.map_const', Multiset.sum_replicate, smul_eq_mul, - mul_one] - · exact Multiset.forall_mem_map_iff.2 fun a _ => monic_X_sub_C a - -theorem units_coeff_zero_smul (c : R[X]ˣ) (p : R[X]) : (c : R[X]).coeff 0 • p = c * p := by - rw [← Polynomial.C_mul', ← Polynomial.eq_C_of_degree_eq_zero (degree_coe_units c)] - -theorem comp_eq_zero_iff : p.comp q = 0 ↔ p = 0 ∨ p.eval (q.coeff 0) = 0 ∧ q = C (q.coeff 0) := by - constructor - · intro h - have key : p.natDegree = 0 ∨ q.natDegree = 0 := by - rw [← mul_eq_zero, ← natDegree_comp, h, natDegree_zero] - replace key := Or.imp eq_C_of_natDegree_eq_zero eq_C_of_natDegree_eq_zero key - cases' key with key key - · rw [key, C_comp] at h - exact Or.inl (key.trans h) - · rw [key, comp_C, C_eq_zero] at h - exact Or.inr ⟨h, key⟩ - · exact fun h => - Or.rec (fun h => by rw [h, zero_comp]) (fun h => by rw [h.2, comp_C, h.1, C_0]) h - lemma aeval_ne_zero_of_isCoprime {R} [CommSemiring R] [Nontrivial S] [Semiring S] [Algebra R S] {p q : R[X]} (h : IsCoprime p q) (s : S) : aeval s p ≠ 0 ∨ aeval s q ≠ 0 := by by_contra! hpq @@ -796,53 +282,4 @@ decreasing_by { end CommRing -section - -variable [Semiring R] [CommRing S] [IsDomain S] (φ : R →+* S) - -theorem isUnit_of_isUnit_leadingCoeff_of_isUnit_map {f : R[X]} (hf : IsUnit f.leadingCoeff) - (H : IsUnit (map φ f)) : IsUnit f := by - have dz := degree_eq_zero_of_isUnit H - rw [degree_map_eq_of_leadingCoeff_ne_zero] at dz - · rw [eq_C_of_degree_eq_zero dz] - refine IsUnit.map C ?_ - convert hf - change coeff f 0 = coeff f (natDegree f) - rw [(degree_eq_iff_natDegree_eq _).1 dz] - · rfl - rintro rfl - simp at H - · intro h - have u : IsUnit (φ f.leadingCoeff) := IsUnit.map φ hf - rw [h] at u - simp at u - -end - -section - -variable [CommRing R] [IsDomain R] [CommRing S] [IsDomain S] (φ : R →+* S) - -/-- A polynomial over an integral domain `R` is irreducible if it is monic and - irreducible after mapping into an integral domain `S`. - -A special case of this lemma is that a polynomial over `ℤ` is irreducible if - it is monic and irreducible over `ℤ/pℤ` for some prime `p`. --/ -theorem Monic.irreducible_of_irreducible_map (f : R[X]) (h_mon : Monic f) - (h_irr : Irreducible (Polynomial.map φ f)) : Irreducible f := by - refine ⟨h_irr.not_unit ∘ IsUnit.map (mapRingHom φ), fun a b h => ?_⟩ - dsimp [Monic] at h_mon - have q := (leadingCoeff_mul a b).symm - rw [← h, h_mon] at q - refine (h_irr.isUnit_or_isUnit <| - (congr_arg (Polynomial.map φ) h).trans (Polynomial.map_mul φ)).imp ?_ ?_ <;> - apply isUnit_of_isUnit_leadingCoeff_of_isUnit_map <;> - apply isUnit_of_mul_eq_one - · exact q - · rw [mul_comm] - exact q - -end - end Polynomial diff --git a/Mathlib/Algebra/Polynomial/Roots.lean b/Mathlib/Algebra/Polynomial/Roots.lean index 4fed6115261d6..3db1a24da52f9 100644 --- a/Mathlib/Algebra/Polynomial/Roots.lean +++ b/Mathlib/Algebra/Polynomial/Roots.lean @@ -3,7 +3,7 @@ Copyright (c) 2018 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Johannes Hölzl, Kim Morrison, Jens Wagemaker, Johan Commelin -/ - +import Mathlib.Algebra.Polynomial.BigOperators import Mathlib.Algebra.Polynomial.RingDivision import Mathlib.RingTheory.Localization.FractionRing import Mathlib.SetTheory.Cardinal.Basic diff --git a/Mathlib/Algebra/Polynomial/SumIteratedDerivative.lean b/Mathlib/Algebra/Polynomial/SumIteratedDerivative.lean index 64723a4d2758b..09b26c906942b 100644 --- a/Mathlib/Algebra/Polynomial/SumIteratedDerivative.lean +++ b/Mathlib/Algebra/Polynomial/SumIteratedDerivative.lean @@ -7,7 +7,6 @@ import Mathlib.Algebra.Polynomial.AlgebraMap import Mathlib.Algebra.Polynomial.BigOperators import Mathlib.Algebra.Polynomial.Degree.Lemmas import Mathlib.Algebra.Polynomial.Derivative -import Mathlib.Algebra.Polynomial.RingDivision /-! # Sum of iterated derivatives diff --git a/Mathlib/Algebra/Polynomial/UnitTrinomial.lean b/Mathlib/Algebra/Polynomial/UnitTrinomial.lean index 10f8454bbd41f..b9f7798bd6c4c 100644 --- a/Mathlib/Algebra/Polynomial/UnitTrinomial.lean +++ b/Mathlib/Algebra/Polynomial/UnitTrinomial.lean @@ -5,6 +5,7 @@ Authors: Thomas Browning -/ import Mathlib.Algebra.Polynomial.Mirror import Mathlib.Data.Int.Order.Units +import Mathlib.RingTheory.Coprime.Basic /-! # Unit Trinomials diff --git a/Mathlib/Analysis/SpecificLimits/Normed.lean b/Mathlib/Analysis/SpecificLimits/Normed.lean index 0e57727abb201..e4f0d3a9330f3 100644 --- a/Mathlib/Analysis/SpecificLimits/Normed.lean +++ b/Mathlib/Analysis/SpecificLimits/Normed.lean @@ -5,6 +5,7 @@ Authors: Anatole Dedecker, Sébastien Gouëzel, Yury Kudryashov, Dylan MacKenzie -/ import Mathlib.Algebra.BigOperators.Module import Mathlib.Algebra.Order.Field.Power +import Mathlib.Algebra.Polynomial.Monic import Mathlib.Analysis.Asymptotics.Asymptotics import Mathlib.Analysis.Normed.Field.InfiniteSum import Mathlib.Analysis.Normed.Module.Basic diff --git a/Mathlib/FieldTheory/RatFunc/Defs.lean b/Mathlib/FieldTheory/RatFunc/Defs.lean index da34170f5b492..c8a7d4b92a7b4 100644 --- a/Mathlib/FieldTheory/RatFunc/Defs.lean +++ b/Mathlib/FieldTheory/RatFunc/Defs.lean @@ -3,8 +3,8 @@ Copyright (c) 2021 Anne Baanen. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Anne Baanen -/ +import Mathlib.Algebra.Polynomial.Degree.Definitions import Mathlib.RingTheory.Localization.FractionRing -import Mathlib.Algebra.Polynomial.RingDivision /-! # The field of rational functions diff --git a/Mathlib/RingTheory/Polynomial/IrreducibleRing.lean b/Mathlib/RingTheory/Polynomial/IrreducibleRing.lean index c2e72452a8f8e..fde8375d8c35a 100644 --- a/Mathlib/RingTheory/Polynomial/IrreducibleRing.lean +++ b/Mathlib/RingTheory/Polynomial/IrreducibleRing.lean @@ -3,7 +3,6 @@ Copyright (c) 2024 Jz Pan. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jz Pan -/ -import Mathlib.Algebra.Polynomial.RingDivision import Mathlib.RingTheory.Polynomial.Nilpotent /-! diff --git a/Mathlib/RingTheory/Polynomial/Pochhammer.lean b/Mathlib/RingTheory/Polynomial/Pochhammer.lean index fcb2365cac7b2..a3dc6084fd575 100644 --- a/Mathlib/RingTheory/Polynomial/Pochhammer.lean +++ b/Mathlib/RingTheory/Polynomial/Pochhammer.lean @@ -3,8 +3,9 @@ Copyright (c) 2020 Kim Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison -/ +import Mathlib.Algebra.Algebra.Basic import Mathlib.Algebra.CharP.Defs -import Mathlib.Algebra.Polynomial.RingDivision +import Mathlib.Algebra.Polynomial.Degree.Lemmas import Mathlib.Tactic.Abel /-! diff --git a/scripts/no_lints_prime_decls.txt b/scripts/no_lints_prime_decls.txt index e1817a2092239..b232c847b8a7a 100644 --- a/scripts/no_lints_prime_decls.txt +++ b/scripts/no_lints_prime_decls.txt @@ -3840,7 +3840,6 @@ Polynomial.natTrailingDegree_eq_support_min' Polynomial.natTrailingDegree_mul' Polynomial.neg' Polynomial.ringHom_ext' -Polynomial.rootMultiplicity_eq_natTrailingDegree' Polynomial.rootMultiplicity_mul' Polynomial.rootMultiplicity_pos' Polynomial.rootSet_maps_to' From 9a0d45feb428a3fe72976bd2bf1193d6d7009692 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 22 Oct 2024 13:55:28 +0000 Subject: [PATCH 273/505] chore(Combinatorics): use newly introduced finset notation (#18055) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit For `s : Finset α`, replace `s.card` with `#s`, `s.filter p` with `{x ∈ s | p x}`, `∑ x ∈ s.filter p, f x` with `∑ x ∈ s with p x, f x`. Rewrap lines around and open the `Finset` locale where necessary. --- .../Additive/AP/Three/Behrend.lean | 11 +- .../Combinatorics/Additive/AP/Three/Defs.lean | 22 +-- .../Combinatorics/Additive/Corner/Roth.lean | 22 +-- .../Combinatorics/Additive/DoublingConst.lean | 48 +++--- .../Additive/ErdosGinzburgZiv.lean | 32 ++-- Mathlib/Combinatorics/Colex.lean | 25 ++- Mathlib/Combinatorics/Configuration.lean | 14 +- .../Combinatorics/Enumerative/Catalan.lean | 2 +- .../Enumerative/DoubleCounting.lean | 60 +++---- Mathlib/Combinatorics/Hall/Basic.lean | 18 +- Mathlib/Combinatorics/Hall/Finite.lean | 44 ++--- Mathlib/Combinatorics/Pigeonhole.lean | 97 ++++++----- Mathlib/Combinatorics/Schnirelmann.lean | 22 +-- .../SetFamily/AhlswedeZhang.lean | 39 +++-- .../SetFamily/CauchyDavenport.lean | 35 ++-- .../SetFamily/Compression/Down.lean | 22 +-- .../SetFamily/Compression/UV.lean | 16 +- .../SetFamily/FourFunctions.lean | 10 +- .../SetFamily/HarrisKleitman.lean | 16 +- .../Combinatorics/SetFamily/Intersecting.lean | 6 +- Mathlib/Combinatorics/SetFamily/Kleitman.lean | 6 +- .../SetFamily/KruskalKatona.lean | 38 ++--- Mathlib/Combinatorics/SetFamily/LYM.lean | 16 +- Mathlib/Combinatorics/SetFamily/Shadow.lean | 24 ++- Mathlib/Combinatorics/SetFamily/Shatter.lean | 14 +- Mathlib/Combinatorics/SimpleGraph/Basic.lean | 2 +- Mathlib/Combinatorics/SimpleGraph/Clique.lean | 19 +-- .../Connectivity/WalkCounting.lean | 9 +- .../Combinatorics/SimpleGraph/DegreeSum.lean | 23 ++- .../Combinatorics/SimpleGraph/Density.lean | 50 +++--- Mathlib/Combinatorics/SimpleGraph/Finite.lean | 29 ++-- .../Combinatorics/SimpleGraph/Operations.lean | 8 +- .../SimpleGraph/Regularity/Bound.lean | 62 ++++--- .../SimpleGraph/Regularity/Chunk.lean | 158 +++++++++--------- .../SimpleGraph/Regularity/Energy.lean | 8 +- .../SimpleGraph/Regularity/Equitabilise.lean | 42 +++-- .../SimpleGraph/Regularity/Increment.lean | 34 ++-- .../SimpleGraph/Regularity/Lemma.lean | 14 +- .../SimpleGraph/Regularity/Uniform.lean | 63 ++++--- .../SimpleGraph/StronglyRegular.lean | 12 +- .../SimpleGraph/Triangle/Basic.lean | 30 ++-- .../SimpleGraph/Triangle/Counting.lean | 34 ++-- .../SimpleGraph/Triangle/Removal.lean | 38 ++--- .../SimpleGraph/Triangle/Tripartite.lean | 4 +- Mathlib/Combinatorics/SimpleGraph/Turan.lean | 20 +-- 45 files changed, 638 insertions(+), 680 deletions(-) diff --git a/Mathlib/Combinatorics/Additive/AP/Three/Behrend.lean b/Mathlib/Combinatorics/Additive/AP/Three/Behrend.lean index 14728c20ba579..54cfd59538a92 100644 --- a/Mathlib/Combinatorics/Additive/AP/Three/Behrend.lean +++ b/Mathlib/Combinatorics/Additive/AP/Three/Behrend.lean @@ -92,15 +92,14 @@ def box (n d : ℕ) : Finset (Fin n → ℕ) := theorem mem_box : x ∈ box n d ↔ ∀ i, x i < d := by simp only [box, Fintype.mem_piFinset, mem_range] @[simp] -theorem card_box : (box n d).card = d ^ n := by simp [box] +theorem card_box : #(box n d) = d ^ n := by simp [box] @[simp] theorem box_zero : box (n + 1) 0 = ∅ := by simp [box] /-- The intersection of the sphere of radius `√k` with the integer points in the positive quadrant. -/ -def sphere (n d k : ℕ) : Finset (Fin n → ℕ) := - (box n d).filter fun x => ∑ i, x i ^ 2 = k +def sphere (n d k : ℕ) : Finset (Fin n → ℕ) := {x ∈ box n d | ∑ i, x i ^ 2 = k} theorem sphere_zero_subset : sphere n d 0 ⊆ 0 := fun x => by simp [sphere, funext_iff] @@ -204,7 +203,7 @@ theorem sum_lt : (∑ i : Fin n, d * (2 * d + 1) ^ (i : ℕ)) < (2 * d + 1) ^ n sum_eq.trans_lt <| (Nat.div_le_self _ 2).trans_lt <| pred_lt (pow_pos (succ_pos _) _).ne' theorem card_sphere_le_rothNumberNat (n d k : ℕ) : - (sphere n d k).card ≤ rothNumberNat ((2 * d - 1) ^ n) := by + #(sphere n d k) ≤ rothNumberNat ((2 * d - 1) ^ n) := by cases n · dsimp; refine (card_le_univ _).trans_eq ?_; rfl cases d @@ -229,7 +228,7 @@ that we then optimize by tweaking the parameters. The (almost) optimal parameter theorem exists_large_sphere_aux (n d : ℕ) : ∃ k ∈ range (n * (d - 1) ^ 2 + 1), - (↑(d ^ n) / ((n * (d - 1) ^ 2 :) + 1) : ℝ) ≤ (sphere n d k).card := by + (↑(d ^ n) / ((n * (d - 1) ^ 2 :) + 1) : ℝ) ≤ #(sphere n d k) := by refine exists_le_card_fiber_of_nsmul_le_card_of_maps_to (fun x hx => ?_) nonempty_range_succ ?_ · rw [mem_range, Nat.lt_succ_iff] exact sum_sq_le_of_mem_box hx @@ -238,7 +237,7 @@ theorem exists_large_sphere_aux (n d : ℕ) : ∃ k ∈ range (n * (d - 1) ^ 2 + exact (cast_add_one_pos _).ne' theorem exists_large_sphere (n d : ℕ) : - ∃ k, ((d ^ n :) / (n * d ^ 2 :) : ℝ) ≤ (sphere n d k).card := by + ∃ k, ((d ^ n :) / (n * d ^ 2 :) : ℝ) ≤ #(sphere n d k) := by obtain ⟨k, -, hk⟩ := exists_large_sphere_aux n d refine ⟨k, ?_⟩ obtain rfl | hn := n.eq_zero_or_pos diff --git a/Mathlib/Combinatorics/Additive/AP/Three/Defs.lean b/Mathlib/Combinatorics/Additive/AP/Three/Defs.lean index 213220d6ea296..f374375f2d4a6 100644 --- a/Mathlib/Combinatorics/Additive/AP/Three/Defs.lean +++ b/Mathlib/Combinatorics/Additive/AP/Three/Defs.lean @@ -253,31 +253,31 @@ subset. The usual Roth number corresponds to `addRothNumber (Finset.range n)`, see `rothNumberNat`."] def mulRothNumber : Finset α →o ℕ := - ⟨fun s ↦ Nat.findGreatest (fun m ↦ ∃ t ⊆ s, t.card = m ∧ ThreeGPFree (t : Set α)) s.card, by + ⟨fun s ↦ Nat.findGreatest (fun m ↦ ∃ t ⊆ s, #t = m ∧ ThreeGPFree (t : Set α)) #s, by rintro t u htu refine Nat.findGreatest_mono (fun m => ?_) (card_le_card htu) rintro ⟨v, hvt, hv⟩ exact ⟨v, hvt.trans htu, hv⟩⟩ @[to_additive] -theorem mulRothNumber_le : mulRothNumber s ≤ s.card := Nat.findGreatest_le s.card +theorem mulRothNumber_le : mulRothNumber s ≤ #s := Nat.findGreatest_le #s @[to_additive] theorem mulRothNumber_spec : - ∃ t ⊆ s, t.card = mulRothNumber s ∧ ThreeGPFree (t : Set α) := - Nat.findGreatest_spec (P := fun m ↦ ∃ t ⊆ s, t.card = m ∧ ThreeGPFree (t : Set α)) + ∃ t ⊆ s, #t = mulRothNumber s ∧ ThreeGPFree (t : Set α) := + Nat.findGreatest_spec (P := fun m ↦ ∃ t ⊆ s, #t = m ∧ ThreeGPFree (t : Set α)) (Nat.zero_le _) ⟨∅, empty_subset _, card_empty, by norm_cast; exact threeGPFree_empty⟩ variable {s t} {n : ℕ} @[to_additive] theorem ThreeGPFree.le_mulRothNumber (hs : ThreeGPFree (s : Set α)) (h : s ⊆ t) : - s.card ≤ mulRothNumber t := + #s ≤ mulRothNumber t := Nat.le_findGreatest (card_le_card h) ⟨s, h, rfl, hs⟩ @[to_additive] theorem ThreeGPFree.mulRothNumber_eq (hs : ThreeGPFree (s : Set α)) : - mulRothNumber s = s.card := + mulRothNumber s = #s := (mulRothNumber_le _).antisymm <| hs.le_mulRothNumber <| Subset.refl _ @[to_additive (attr := simp)] @@ -295,9 +295,9 @@ theorem mulRothNumber_union_le (s t : Finset α) : mulRothNumber (s ∪ t) ≤ mulRothNumber s + mulRothNumber t := let ⟨u, hus, hcard, hu⟩ := mulRothNumber_spec (s ∪ t) calc - mulRothNumber (s ∪ t) = u.card := hcard.symm - _ = (u ∩ s ∪ u ∩ t).card := by rw [← inter_union_distrib_left, inter_eq_left.2 hus] - _ ≤ (u ∩ s).card + (u ∩ t).card := card_union_le _ _ + mulRothNumber (s ∪ t) = #u := hcard.symm + _ = #(u ∩ s ∪ u ∩ t) := by rw [← inter_union_distrib_left, inter_eq_left.2 hus] + _ ≤ #(u ∩ s) + #(u ∩ t) := card_union_le _ _ _ ≤ mulRothNumber s + mulRothNumber t := _root_.add_le_add ((hu.mono inter_subset_left).le_mulRothNumber inter_subset_right) ((hu.mono inter_subset_left).le_mulRothNumber inter_subset_right) @@ -407,13 +407,13 @@ theorem rothNumberNat_le (N : ℕ) : rothNumberNat N ≤ N := (addRothNumber_le _).trans (card_range _).le theorem rothNumberNat_spec (n : ℕ) : - ∃ t ⊆ range n, t.card = rothNumberNat n ∧ ThreeAPFree (t : Set ℕ) := + ∃ t ⊆ range n, #t = rothNumberNat n ∧ ThreeAPFree (t : Set ℕ) := addRothNumber_spec _ /-- A verbose specialization of `threeAPFree.le_addRothNumber`, sometimes convenient in practice. -/ theorem ThreeAPFree.le_rothNumberNat (s : Finset ℕ) (hs : ThreeAPFree (s : Set ℕ)) - (hsn : ∀ x ∈ s, x < n) (hsk : s.card = k) : k ≤ rothNumberNat n := + (hsn : ∀ x ∈ s, x < n) (hsk : #s = k) : k ≤ rothNumberNat n := hsk.ge.trans <| hs.le_addRothNumber fun x hx => mem_range.2 <| hsn x hx /-- The Roth number is a subadditive function. Note that by Fekete's lemma this shows that diff --git a/Mathlib/Combinatorics/Additive/Corner/Roth.lean b/Mathlib/Combinatorics/Additive/Corner/Roth.lean index 73840313536fa..75cfac6bab75d 100644 --- a/Mathlib/Combinatorics/Additive/Corner/Roth.lean +++ b/Mathlib/Combinatorics/Additive/Corner/Roth.lean @@ -40,7 +40,7 @@ private lemma mk_mem_triangleIndices : (a, b, c) ∈ triangleIndices A ↔ (a, b rintro ⟨_, _, h₁, rfl, rfl, h₂⟩ exact ⟨h₁, h₂⟩ -@[simp] private lemma card_triangleIndices : (triangleIndices A).card = A.card := card_map _ +@[simp] private lemma card_triangleIndices : #(triangleIndices A) = #A := card_map _ private instance triangleIndices.instExplicitDisjoint : ExplicitDisjoint (triangleIndices A) := by constructor @@ -55,7 +55,7 @@ private lemma noAccidental (hs : IsCornerFree (A : Set (G × G))) : simp only [mk_mem_triangleIndices] at ha hb hc exact .inl <| hs ⟨hc.1, hb.1, ha.1, hb.2.symm.trans ha.2⟩ -private lemma farFromTriangleFree_graph [Fintype G] [DecidableEq G] (hε : ε * card G ^ 2 ≤ A.card) : +private lemma farFromTriangleFree_graph [Fintype G] [DecidableEq G] (hε : ε * card G ^ 2 ≤ #A) : (graph <| triangleIndices A).FarFromTriangleFree (ε / 9) := by refine farFromTriangleFree _ ?_ simp_rw [card_triangleIndices, mul_comm_div, Nat.cast_pow, Nat.cast_add] @@ -79,7 +79,7 @@ noncomputable def cornersTheoremBound (ε : ℝ) : ℕ := ⌊(triangleRemovalBou The maximum density of a corner-free set in `G × G` goes to zero as `|G|` tends to infinity. -/ theorem corners_theorem (ε : ℝ) (hε : 0 < ε) (hG : cornersTheoremBound ε ≤ card G) - (A : Finset (G × G)) (hAε : ε * card G ^ 2 ≤ A.card) : ¬ IsCornerFree (A : Set (G × G)) := by + (A : Finset (G × G)) (hAε : ε * card G ^ 2 ≤ #A) : ¬ IsCornerFree (A : Set (G × G)) := by rintro hA rw [cornersTheoremBound, Nat.add_one_le_iff] at hG have hε₁ : ε ≤ 1 := by @@ -103,7 +103,7 @@ theorem corners_theorem (ε : ℝ) (hε : 0 < ε) (hG : cornersTheoremBound ε The maximum density of a corner-free set in `{1, ..., n} × {1, ..., n}` goes to zero as `n` tends to infinity. -/ theorem corners_theorem_nat (hε : 0 < ε) (hn : cornersTheoremBound (ε / 9) ≤ n) - (A : Finset (ℕ × ℕ)) (hAn : A ⊆ range n ×ˢ range n) (hAε : ε * n ^ 2 ≤ A.card) : + (A : Finset (ℕ × ℕ)) (hAn : A ⊆ range n ×ˢ range n) (hAε : ε * n ^ 2 ≤ #A) : ¬ IsCornerFree (A : Set (ℕ × ℕ)) := by rintro hA rw [← coe_subset, coe_product] at hAn @@ -128,7 +128,7 @@ theorem corners_theorem_nat (hε : 0 < ε) (hn : cornersTheoremBound (ε / 9) _ = ε / 9 * (2 * n + 1) ^ 2 := by simp _ ≤ ε / 9 * (2 * n + n) ^ 2 := by gcongr; simp; unfold cornersTheoremBound at hn; omega _ = ε * n ^ 2 := by ring - _ ≤ A.card := hAε + _ ≤ #A := hAε _ = _ := by rw [card_image_of_injOn] have : Set.InjOn Nat.cast (range n) := @@ -139,15 +139,15 @@ theorem corners_theorem_nat (hε : 0 < ε) (hn : cornersTheoremBound (ε / 9) The maximum density of a 3AP-free set in `G` goes to zero as `|G|` tends to infinity. -/ theorem roth_3ap_theorem (ε : ℝ) (hε : 0 < ε) (hG : cornersTheoremBound ε ≤ card G) - (A : Finset G) (hAε : ε * card G ≤ A.card) : ¬ ThreeAPFree (A : Set G) := by + (A : Finset G) (hAε : ε * card G ≤ #A) : ¬ ThreeAPFree (A : Set G) := by rintro hA classical let B : Finset (G × G) := univ.filter fun (x, y) ↦ y - x ∈ A - have : ε * card G ^ 2 ≤ B.card := by + have : ε * card G ^ 2 ≤ #B := by calc _ = card G * (ε * card G) := by ring - _ ≤ card G * A.card := by gcongr - _ = B.card := ?_ + _ ≤ card G * #A := by gcongr + _ = #B := ?_ norm_cast rw [← card_univ, ← card_product] exact card_equiv ((Equiv.refl _).prodShear fun a ↦ Equiv.addLeft a) (by simp [B]) @@ -164,7 +164,7 @@ theorem roth_3ap_theorem (ε : ℝ) (hε : 0 < ε) (hG : cornersTheoremBound ε The maximum density of a 3AP-free set in `{1, ..., n}` goes to zero as `n` tends to infinity. -/ theorem roth_3ap_theorem_nat (ε : ℝ) (hε : 0 < ε) (hG : cornersTheoremBound (ε / 3) ≤ n) - (A : Finset ℕ) (hAn : A ⊆ range n) (hAε : ε * n ≤ A.card) : ¬ ThreeAPFree (A : Set ℕ) := by + (A : Finset ℕ) (hAn : A ⊆ range n) (hAε : ε * n ≤ #A) : ¬ ThreeAPFree (A : Set ℕ) := by rintro hA rw [← coe_subset, coe_range] at hAn have : A = Fin.val '' (Nat.cast '' A : Set (Fin (2 * n).succ)) := by @@ -185,7 +185,7 @@ theorem roth_3ap_theorem_nat (ε : ℝ) (hε : 0 < ε) (hG : cornersTheoremBound _ = ε / 3 * (2 * n + 1) := by simp _ ≤ ε / 3 * (2 * n + n) := by gcongr; simp; unfold cornersTheoremBound at hG; omega _ = ε * n := by ring - _ ≤ A.card := hAε + _ ≤ #A := hAε _ = _ := by rw [card_image_of_injOn] exact (CharP.natCast_injOn_Iio (Fin (2 * n).succ) (2 * n).succ).mono <| hAn.trans <| by diff --git a/Mathlib/Combinatorics/Additive/DoublingConst.lean b/Mathlib/Combinatorics/Additive/DoublingConst.lean index f9f65e418d246..c4534bdd4b0d4 100644 --- a/Mathlib/Combinatorics/Additive/DoublingConst.lean +++ b/Mathlib/Combinatorics/Additive/DoublingConst.lean @@ -30,7 +30,7 @@ The notation `σₘ[A, B]` is available in scope `Combinatorics.Additive`. -/ "The doubling constant `σ[A, B]` of two finsets `A` and `B` in a group is `|A + B| / |A|`. The notation `σ[A, B]` is available in scope `Combinatorics.Additive`."] -def mulConst (A B : Finset G) : ℚ≥0 := (A * B).card / A.card +def mulConst (A B : Finset G) : ℚ≥0 := #(A * B) / #A /-- The difference constant `δₘ[A, B]` of two finsets `A` and `B` in a group is `|A / B| / |A|`. @@ -39,7 +39,7 @@ The notation `δₘ[A, B]` is available in scope `Combinatorics.Additive`. -/ "The difference constant `σ[A, B]` of two finsets `A` and `B` in a group is `|A - B| / |A|`. The notation `δ[A, B]` is available in scope `Combinatorics.Additive`."] -def divConst (A B : Finset G) : ℚ≥0 := (A / B).card / A.card +def divConst (A B : Finset G) : ℚ≥0 := #(A / B) / #A /-- The doubling constant `σₘ[A, B]` of two finsets `A` and `B` in a group is `|A * B| / |A|`. -/ scoped[Combinatorics.Additive] notation3:max "σₘ[" A ", " B "]" => Finset.mulConst A B @@ -68,23 +68,23 @@ scoped[Combinatorics.Additive] notation3:max "δ[" A "]" => Finset.subConst A A open scoped Combinatorics.Additive @[to_additive (attr := simp) addConst_mul_card] -lemma mulConst_mul_card (A B : Finset G) : σₘ[A, B] * A.card = (A * B).card := by +lemma mulConst_mul_card (A B : Finset G) : σₘ[A, B] * #A = #(A * B) := by obtain rfl | hA := A.eq_empty_or_nonempty · simp · exact div_mul_cancel₀ _ (by positivity) @[to_additive (attr := simp) subConst_mul_card] -lemma divConst_mul_card (A B : Finset G) : δₘ[A, B] * A.card = (A / B).card := by +lemma divConst_mul_card (A B : Finset G) : δₘ[A, B] * #A = #(A / B) := by obtain rfl | hA := A.eq_empty_or_nonempty · simp · exact div_mul_cancel₀ _ (by positivity) @[to_additive (attr := simp) card_mul_addConst] -lemma card_mul_mulConst (A B : Finset G) : A.card * σₘ[A, B] = (A * B).card := by +lemma card_mul_mulConst (A B : Finset G) : #A * σₘ[A, B] = #(A * B) := by rw [mul_comm, mulConst_mul_card] @[to_additive (attr := simp) card_mul_subConst] -lemma card_mul_divConst (A B : Finset G) : A.card * δₘ[A, B] = (A / B).card := by +lemma card_mul_divConst (A B : Finset G) : #A * δₘ[A, B] = #(A / B) := by rw [mul_comm, divConst_mul_card] @[to_additive (attr := simp)] @@ -124,44 +124,44 @@ end Fintype variable {𝕜 : Type*} [Semifield 𝕜] [CharZero 𝕜] -lemma cast_addConst (A B : Finset G') : (σ[A, B] : 𝕜) = (A + B).card / A.card := by +lemma cast_addConst (A B : Finset G') : (σ[A, B] : 𝕜) = #(A + B) / #A := by simp [addConst] -lemma cast_subConst (A B : Finset G') : (δ[A, B] : 𝕜) = (A - B).card / A.card := by +lemma cast_subConst (A B : Finset G') : (δ[A, B] : 𝕜) = #(A - B) / #A := by simp [subConst] @[to_additive existing] -lemma cast_mulConst (A B : Finset G) : (σₘ[A, B] : 𝕜) = (A * B).card / A.card := by simp [mulConst] +lemma cast_mulConst (A B : Finset G) : (σₘ[A, B] : 𝕜) = #(A * B) / #A := by simp [mulConst] @[to_additive existing] -lemma cast_divConst (A B : Finset G) : (δₘ[A, B] : 𝕜) = (A / B).card / A.card := by simp [divConst] +lemma cast_divConst (A B : Finset G) : (δₘ[A, B] : 𝕜) = #(A / B) / #A := by simp [divConst] -lemma cast_addConst_mul_card (A B : Finset G') : (σ[A, B] * A.card : 𝕜) = (A + B).card := by +lemma cast_addConst_mul_card (A B : Finset G') : (σ[A, B] * #A : 𝕜) = #(A + B) := by norm_cast; exact addConst_mul_card _ _ -lemma cast_subConst_mul_card (A B : Finset G') : (δ[A, B] * A.card : 𝕜) = (A - B).card := by +lemma cast_subConst_mul_card (A B : Finset G') : (δ[A, B] * #A : 𝕜) = #(A - B) := by norm_cast; exact subConst_mul_card _ _ -lemma card_mul_cast_addConst (A B : Finset G') : (A.card * σ[A, B] : 𝕜) = (A + B).card := by +lemma card_mul_cast_addConst (A B : Finset G') : (#A * σ[A, B] : 𝕜) = #(A + B) := by norm_cast; exact card_mul_addConst _ _ -lemma card_mul_cast_subConst (A B : Finset G') : (A.card * δ[A, B] : 𝕜) = (A - B).card := by +lemma card_mul_cast_subConst (A B : Finset G') : (#A * δ[A, B] : 𝕜) = #(A - B) := by norm_cast; exact card_mul_subConst _ _ @[to_additive (attr := simp) existing cast_addConst_mul_card] -lemma cast_mulConst_mul_card (A B : Finset G) : (σₘ[A, B] * A.card : 𝕜) = (A * B).card := by +lemma cast_mulConst_mul_card (A B : Finset G) : (σₘ[A, B] * #A : 𝕜) = #(A * B) := by norm_cast; exact mulConst_mul_card _ _ @[to_additive (attr := simp) existing cast_subConst_mul_card] -lemma cast_divConst_mul_card (A B : Finset G) : (δₘ[A, B] * A.card : 𝕜) = (A / B).card := by +lemma cast_divConst_mul_card (A B : Finset G) : (δₘ[A, B] * #A : 𝕜) = #(A / B) := by norm_cast; exact divConst_mul_card _ _ @[to_additive (attr := simp) existing card_mul_cast_addConst] -lemma card_mul_cast_mulConst (A B : Finset G) : (A.card * σₘ[A, B] : 𝕜) = (A * B).card := by +lemma card_mul_cast_mulConst (A B : Finset G) : (#A * σₘ[A, B] : 𝕜) = #(A * B) := by norm_cast; exact card_mul_mulConst _ _ @[to_additive (attr := simp) existing card_mul_cast_subConst] -lemma card_mul_cast_divConst (A B : Finset G) : (A.card * δₘ[A, B] : 𝕜) = (A / B).card := by +lemma card_mul_cast_divConst (A B : Finset G) : (#A * δₘ[A, B] : 𝕜) = #(A / B) := by norm_cast; exact card_mul_divConst _ _ end Group @@ -189,10 +189,10 @@ This is a consequence of the Ruzsa triangle inequality."] lemma mulConst_le_divConst_sq : σₘ[A] ≤ δₘ[A] ^ 2 := by obtain rfl | hA' := A.eq_empty_or_nonempty · simp - refine le_of_mul_le_mul_right ?_ (by positivity : (0 : ℚ≥0) < A.card * A.card) + refine le_of_mul_le_mul_right ?_ (by positivity : (0 : ℚ≥0) < #A * #A) calc - _ = (A * A).card * (A.card : ℚ≥0) := by rw [← mul_assoc, mulConst_mul_card] - _ ≤ (A / A).card * (A / A).card := by norm_cast; exact ruzsa_triangle_inequality_mul_div_div .. + _ = #(A * A) * (#A : ℚ≥0) := by rw [← mul_assoc, mulConst_mul_card] + _ ≤ #(A / A) * #(A / A) := by norm_cast; exact ruzsa_triangle_inequality_mul_div_div .. _ = _ := by rw [← divConst_mul_card]; ring /-- If `A` has small doubling, then it has small difference, with the constant squared. @@ -205,10 +205,10 @@ This is a consequence of the Ruzsa triangle inequality."] lemma divConst_le_mulConst_sq : δₘ[A] ≤ σₘ[A] ^ 2 := by obtain rfl | hA' := A.eq_empty_or_nonempty · simp - refine le_of_mul_le_mul_right ?_ (by positivity : (0 : ℚ≥0) < A.card * A.card) + refine le_of_mul_le_mul_right ?_ (by positivity : (0 : ℚ≥0) < #A * #A) calc - _ = (A / A).card * (A.card : ℚ≥0) := by rw [← mul_assoc, divConst_mul_card] - _ ≤ (A * A).card * (A * A).card := by norm_cast; exact ruzsa_triangle_inequality_div_mul_mul .. + _ = #(A / A) * (#A : ℚ≥0) := by rw [← mul_assoc, divConst_mul_card] + _ ≤ #(A * A) * #(A * A) := by norm_cast; exact ruzsa_triangle_inequality_div_mul_mul .. _ = _ := by rw [← mulConst_mul_card]; ring end CommGroup diff --git a/Mathlib/Combinatorics/Additive/ErdosGinzburgZiv.lean b/Mathlib/Combinatorics/Additive/ErdosGinzburgZiv.lean index 164ae582fd759..f20bc9174d851 100644 --- a/Mathlib/Combinatorics/Additive/ErdosGinzburgZiv.lean +++ b/Mathlib/Combinatorics/Additive/ErdosGinzburgZiv.lean @@ -50,8 +50,8 @@ private lemma totalDegree_f₁_add_totalDegree_f₂ {a : ι → ZMod p} : Any sequence of `2 * p - 1` elements of `ZMod p` contains a subsequence of `p` elements whose sum is zero. -/ -private theorem ZMod.erdos_ginzburg_ziv_prime (a : ι → ZMod p) (hs : s.card = 2 * p - 1) : - ∃ t ⊆ s, t.card = p ∧ ∑ i ∈ t, a i = 0 := by +private theorem ZMod.erdos_ginzburg_ziv_prime (a : ι → ZMod p) (hs : #s = 2 * p - 1) : + ∃ t ⊆ s, #t = p ∧ ∑ i ∈ t, a i = 0 := by haveI : NeZero p := inferInstance classical -- Let `N` be the number of common roots of our polynomials `f₁` and `f₂` (`f s ff` and `f s tt`). @@ -69,7 +69,7 @@ private theorem ZMod.erdos_ginzburg_ziv_prime (a : ι → ZMod p) (hs : s.card = obtain ⟨x, hx⟩ := Fintype.exists_ne_of_one_lt_card ((Fact.out : p.Prime).one_lt.trans_le <| Nat.le_of_dvd hN₀ hpN) zero_sol -- This common root gives us the required subsequence, namely the `i ∈ s` such that `x i ≠ 0`. - refine ⟨(s.attach.filter fun a ↦ x.1 a ≠ 0).map ⟨(↑), Subtype.val_injective⟩, ?_, ?_, ?_⟩ + refine ⟨({a | x.1 a ≠ 0} : Finset _).map ⟨(↑), Subtype.val_injective⟩, ?_, ?_, ?_⟩ · simp (config := { contextual := true }) [subset_iff] -- From `f₁ x = 0`, we get that `p` divides the number of `a` such that `x a ≠ 0`. · rw [card_map] @@ -81,7 +81,7 @@ private theorem ZMod.erdos_ginzburg_ziv_prime (a : ι → ZMod p) (hs : s.card = · rw [← CharP.cast_eq_zero_iff (ZMod p), ← Finset.sum_boole] simpa only [f₁, map_sum, ZMod.pow_card_sub_one, map_pow, eval_X] using x.2.1 -- And it is at most `2 * p - 1`, so it must be `p`. - · rw [Finset.card_attach, hs] + · rw [univ_eq_attach, card_attach, hs] exact tsub_lt_self (mul_pos zero_lt_two (Fact.out : p.Prime).pos) zero_lt_one -- From `f₂ x = 0`, we get that `p` divides the sum of the `a ∈ s` such that `x a ≠ 0`. · simpa [f₂, ZMod.pow_card_sub_one, Finset.sum_filter] using x.2.2 @@ -90,8 +90,8 @@ private theorem ZMod.erdos_ginzburg_ziv_prime (a : ι → ZMod p) (hs : s.card = Any sequence of `2 * p - 1` elements of `ℤ` contains a subsequence of `p` elements whose sum is divisible by `p`. -/ -private theorem Int.erdos_ginzburg_ziv_prime (a : ι → ℤ) (hs : s.card = 2 * p - 1) : - ∃ t ⊆ s, t.card = p ∧ ↑p ∣ ∑ i ∈ t, a i := by +private theorem Int.erdos_ginzburg_ziv_prime (a : ι → ℤ) (hs : #s = 2 * p - 1) : + ∃ t ⊆ s, #t = p ∧ ↑p ∣ ∑ i ∈ t, a i := by simpa [← Int.cast_sum, ZMod.intCast_zmod_eq_zero_iff_dvd] using ZMod.erdos_ginzburg_ziv_prime (Int.cast ∘ a) hs @@ -104,8 +104,8 @@ variable {n : ℕ} {s : Finset ι} Any sequence of at least `2 * n - 1` elements of `ℤ` contains a subsequence of `n` elements whose sum is divisible by `n`. -/ -theorem Int.erdos_ginzburg_ziv (a : ι → ℤ) (hs : 2 * n - 1 ≤ s.card) : - ∃ t ⊆ s, t.card = n ∧ ↑n ∣ ∑ i ∈ t, a i := by +theorem Int.erdos_ginzburg_ziv (a : ι → ℤ) (hs : 2 * n - 1 ≤ #s) : + ∃ t ⊆ s, #t = n ∧ ↑n ∣ ∑ i ∈ t, a i := by classical -- Do induction on the prime factorisation of `n`. Note that we will apply the induction -- hypothesis with `ι := Finset ι`, so we need to generalise. @@ -125,9 +125,9 @@ theorem Int.erdos_ginzburg_ziv (a : ι → ℤ) (hs : 2 * n - 1 ≤ s.card) : -- these sets whose sum is divisible by `m * n`. | composite m hm ihm n hn ihn => -- First, show that it is enough to have those `2 * m - 1` sets. - suffices ∀ k ≤ 2 * m - 1, ∃ 𝒜 : Finset (Finset ι), 𝒜.card = k ∧ + suffices ∀ k ≤ 2 * m - 1, ∃ 𝒜 : Finset (Finset ι), #𝒜 = k ∧ (𝒜 : Set (Finset ι)).Pairwise _root_.Disjoint ∧ - ∀ ⦃t⦄, t ∈ 𝒜 → t ⊆ s ∧ t.card = n ∧ ↑n ∣ ∑ i ∈ t, a i by + ∀ ⦃t⦄, t ∈ 𝒜 → t ⊆ s ∧ #t = n ∧ ↑n ∣ ∑ i ∈ t, a i by -- Assume `𝒜` is a family of `2 * m - 1` sets, each of size `n` and sum divisible by `n`. obtain ⟨𝒜, h𝒜card, h𝒜disj, h𝒜⟩ := this _ le_rfl -- By induction hypothesis on `m`, find a subfamily `ℬ` of size `m` such that the sum over @@ -150,13 +150,13 @@ theorem Int.erdos_ginzburg_ziv (a : ι → ℤ) (hs : 2 * n - 1 ≤ s.card) : obtain ⟨𝒜, h𝒜card, h𝒜disj, h𝒜⟩ := ih (Nat.le_of_succ_le hk) -- There are at least `2 * (m * n) - 1 - k * n ≥ 2 * m - 1` elements in `s` that have not been -- taken in any element of `𝒜`. - have : 2 * n - 1 ≤ (s \ 𝒜.biUnion id).card := by + have : 2 * n - 1 ≤ #(s \ 𝒜.biUnion id) := by calc _ ≤ (2 * m - k) * n - 1 := by gcongr; omega - _ = (2 * (m * n) - 1) - ∑ t ∈ 𝒜, t.card := by + _ = (2 * (m * n) - 1) - ∑ t ∈ 𝒜, #t := by rw [tsub_mul, mul_assoc, tsub_right_comm, sum_const_nat fun t ht ↦ (h𝒜 ht).2.1, h𝒜card] - _ ≤ s.card - (𝒜.biUnion id).card := by gcongr; exact card_biUnion_le - _ ≤ (s \ 𝒜.biUnion id).card := le_card_sdiff .. + _ ≤ #s - #(𝒜.biUnion id) := by gcongr; exact card_biUnion_le + _ ≤ #(s \ 𝒜.biUnion id) := le_card_sdiff .. -- So by the induction hypothesis on `n` we can find a new set `t` of size `n` and sum divisible -- by `n`. obtain ⟨t₀, ht₀, ht₀card, ht₀sum⟩ := ihn a this @@ -177,8 +177,8 @@ theorem Int.erdos_ginzburg_ziv (a : ι → ℤ) (hs : 2 * n - 1 ≤ s.card) : Any sequence of at least `2 * n - 1` elements of `ZMod n` contains a subsequence of `n` elements whose sum is zero. -/ -theorem ZMod.erdos_ginzburg_ziv (a : ι → ZMod n) (hs : 2 * n - 1 ≤ s.card) : - ∃ t ⊆ s, t.card = n ∧ ∑ i ∈ t, a i = 0 := by +theorem ZMod.erdos_ginzburg_ziv (a : ι → ZMod n) (hs : 2 * n - 1 ≤ #s) : + ∃ t ⊆ s, #t = n ∧ ∑ i ∈ t, a i = 0 := by simpa [← ZMod.intCast_zmod_eq_zero_iff_dvd] using Int.erdos_ginzburg_ziv (ZMod.cast ∘ a) hs /-- The **Erdős–Ginzburg–Ziv theorem** for `ℤ` for multiset. diff --git a/Mathlib/Combinatorics/Colex.lean b/Mathlib/Combinatorics/Colex.lean index 1513265923d82..51ab9b09db4e7 100644 --- a/Mathlib/Combinatorics/Colex.lean +++ b/Mathlib/Combinatorics/Colex.lean @@ -102,7 +102,7 @@ instance instLE : LE (Colex α) where private lemma trans_aux (hst : toColex s ≤ toColex t) (htu : toColex t ≤ toColex u) (has : a ∈ s) (hat : a ∉ t) : ∃ b, b ∈ u ∧ b ∉ s ∧ a ≤ b := by classical - let s' : Finset α := s.filter fun b ↦ b ∉ t ∧ a ≤ b + let s' : Finset α := {b ∈ s | b ∉ t ∧ a ≤ b} have ⟨b, hb, hbmax⟩ := exists_maximal s' ⟨a, by simp [s', has, hat]⟩ simp only [s', mem_filter, and_imp] at hb hbmax have ⟨c, hct, hcs, hbc⟩ := hst hb.1 hb.2.1 @@ -338,10 +338,10 @@ lemma lt_iff_max'_mem {s t : Colex α} : rw [lt_iff_le_and_ne, le_iff_max'_mem]; aesop lemma lt_iff_exists_filter_lt : - toColex s < toColex t ↔ ∃ w ∈ t \ s, s.filter (w < ·) = t.filter (w < ·) := by + toColex s < toColex t ↔ ∃ w ∈ t \ s, {a ∈ s | w < a} = {a ∈ t | w < a} := by simp only [lt_iff_exists_forall_lt, mem_sdiff, filter_inj, and_assoc] refine ⟨fun h ↦ ?_, ?_⟩ - · let u := (t \ s).filter fun w ↦ ∀ a ∈ s, a ∉ t → a < w + · let u := {w ∈ t \ s | ∀ a ∈ s, a ∉ t → a < w} have mem_u {w : α} : w ∈ u ↔ w ∈ t ∧ w ∉ s ∧ ∀ a ∈ s, a ∉ t → a < w := by simp [u, and_assoc] have hu : u.Nonempty := h.imp fun _ ↦ mem_u.2 let m := max' _ hu @@ -355,8 +355,8 @@ lemma lt_iff_exists_filter_lt : by_contra! hwa exact hat <| (hw <| hwa.lt_of_ne <| ne_of_mem_of_not_mem hwt hat).1 has -/-- If `s ≤ t` in colex and `s.card ≤ t.card`, then `s \ {a} ≤ t \ {min t}` for any `a ∈ s`. -/ -lemma erase_le_erase_min' (hst : toColex s ≤ toColex t) (hcard : s.card ≤ t.card) (ha : a ∈ s) : +/-- If `s ≤ t` in colex and `#s ≤ #t`, then `s \ {a} ≤ t \ {min t}` for any `a ∈ s`. -/ +lemma erase_le_erase_min' (hst : toColex s ≤ toColex t) (hcard : #s ≤ #t) (ha : a ∈ s) : toColex (s.erase a) ≤ toColex (t.erase <| min' t <| card_pos.1 <| (card_pos.2 ⟨a, ha⟩).trans_le hcard) := by generalize_proofs ht @@ -434,7 +434,7 @@ end Fintype downwards closed with respect to colex among sets of size `r`. -/ def IsInitSeg (𝒜 : Finset (Finset α)) (r : ℕ) : Prop := (𝒜 : Set (Finset α)).Sized r ∧ - ∀ ⦃s t : Finset α⦄, s ∈ 𝒜 → toColex t < toColex s ∧ t.card = r → t ∈ 𝒜 + ∀ ⦃s t : Finset α⦄, s ∈ 𝒜 → toColex t < toColex s ∧ #t = r → t ∈ 𝒜 @[simp] lemma isInitSeg_empty : IsInitSeg (∅ : Finset (Finset α)) r := by simp [IsInitSeg] @@ -454,24 +454,23 @@ lemma IsInitSeg.total (h₁ : IsInitSeg 𝒜₁ r) (h₂ : IsInitSeg 𝒜₂ r) variable [Fintype α] -/-- The initial segment of the colexicographic order on sets with `s.card` elements and ending at +/-- The initial segment of the colexicographic order on sets with `#s` elements and ending at `s`. -/ -def initSeg (s : Finset α) : Finset (Finset α) := - univ.filter fun t ↦ s.card = t.card ∧ toColex t ≤ toColex s +def initSeg (s : Finset α) : Finset (Finset α) := {t | #s = #t ∧ toColex t ≤ toColex s} @[simp] -lemma mem_initSeg : t ∈ initSeg s ↔ s.card = t.card ∧ toColex t ≤ toColex s := by simp [initSeg] +lemma mem_initSeg : t ∈ initSeg s ↔ #s = #t ∧ toColex t ≤ toColex s := by simp [initSeg] lemma mem_initSeg_self : s ∈ initSeg s := by simp @[simp] lemma initSeg_nonempty : (initSeg s).Nonempty := ⟨s, mem_initSeg_self⟩ -lemma isInitSeg_initSeg : IsInitSeg (initSeg s) s.card := by +lemma isInitSeg_initSeg : IsInitSeg (initSeg s) #s := by refine ⟨fun t ht => (mem_initSeg.1 ht).1.symm, fun t₁ t₂ ht₁ ht₂ ↦ mem_initSeg.2 ⟨ht₂.2.symm, ?_⟩⟩ rw [mem_initSeg] at ht₁ exact ht₂.1.le.trans ht₁.2 lemma IsInitSeg.exists_initSeg (h𝒜 : IsInitSeg 𝒜 r) (h𝒜₀ : 𝒜.Nonempty) : - ∃ s : Finset α, s.card = r ∧ 𝒜 = initSeg s := by + ∃ s : Finset α, #s = r ∧ 𝒜 = initSeg s := by have hs := sup'_mem (ofColex ⁻¹' 𝒜) (LinearOrder.supClosed _) 𝒜 h𝒜₀ toColex (fun a ha ↦ by simpa using ha) refine ⟨_, h𝒜.1 hs, ?_⟩ @@ -487,7 +486,7 @@ lemma IsInitSeg.exists_initSeg (h𝒜 : IsInitSeg 𝒜 r) (h𝒜₀ : 𝒜.Nonem /-- Being a nonempty initial segment of colex is equivalent to being an `initSeg`. -/ lemma isInitSeg_iff_exists_initSeg : - IsInitSeg 𝒜 r ∧ 𝒜.Nonempty ↔ ∃ s : Finset α, s.card = r ∧ 𝒜 = initSeg s := by + IsInitSeg 𝒜 r ∧ 𝒜.Nonempty ↔ ∃ s : Finset α, #s = r ∧ 𝒜 = initSeg s := by refine ⟨fun h𝒜 ↦ h𝒜.1.exists_initSeg h𝒜.2, ?_⟩ rintro ⟨s, rfl, rfl⟩ exact ⟨isInitSeg_initSeg, initSeg_nonempty⟩ diff --git a/Mathlib/Combinatorics/Configuration.lean b/Mathlib/Combinatorics/Configuration.lean index a703f011ed4da..79f8392c7bf56 100644 --- a/Mathlib/Combinatorics/Configuration.lean +++ b/Mathlib/Combinatorics/Configuration.lean @@ -117,27 +117,27 @@ theorem Nondegenerate.exists_injective_of_card_le [Nondegenerate P L] [Fintype P (h : Fintype.card L ≤ Fintype.card P) : ∃ f : L → P, Function.Injective f ∧ ∀ l, f l ∉ l := by classical let t : L → Finset P := fun l => Set.toFinset { p | p ∉ l } - suffices ∀ s : Finset L, s.card ≤ (s.biUnion t).card by + suffices ∀ s : Finset L, #s ≤ (s.biUnion t).card by -- Hall's marriage theorem obtain ⟨f, hf1, hf2⟩ := (Finset.all_card_le_biUnion_card_iff_exists_injective t).mp this exact ⟨f, hf1, fun l => Set.mem_toFinset.mp (hf2 l)⟩ intro s - by_cases hs₀ : s.card = 0 - -- If `s = ∅`, then `s.card = 0 ≤ (s.bUnion t).card` + by_cases hs₀ : #s = 0 + -- If `s = ∅`, then `#s = 0 ≤ #(s.bUnion t)` · simp_rw [hs₀, zero_le] - by_cases hs₁ : s.card = 1 + by_cases hs₁ : #s = 1 -- If `s = {l}`, then pick a point `p ∉ l` · obtain ⟨l, rfl⟩ := Finset.card_eq_one.mp hs₁ obtain ⟨p, hl⟩ := exists_point (P := P) l rw [Finset.card_singleton, Finset.singleton_biUnion, Nat.one_le_iff_ne_zero] exact Finset.card_ne_zero_of_mem (Set.mem_toFinset.mpr hl) - suffices (s.biUnion t)ᶜ.card ≤ sᶜ.card by + suffices #(s.biUnion t)ᶜ ≤ #sᶜ by -- Rephrase in terms of complements (uses `h`) rw [Finset.card_compl, Finset.card_compl, tsub_le_iff_left] at this replace := h.trans this rwa [← add_tsub_assoc_of_le s.card_le_univ, le_tsub_iff_left (le_add_left s.card_le_univ), add_le_add_iff_right] at this - have hs₂ : (s.biUnion t)ᶜ.card ≤ 1 := by + have hs₂ : #(s.biUnion t)ᶜ ≤ 1 := by -- At most one line through two points of `s` refine Finset.card_le_one_iff.mpr @fun p₁ p₂ hp₁ hp₂ => ?_ simp_rw [t, Finset.mem_compl, Finset.mem_biUnion, not_exists, not_and, @@ -145,7 +145,7 @@ theorem Nondegenerate.exists_injective_of_card_le [Nondegenerate P L] [Fintype P obtain ⟨l₁, l₂, hl₁, hl₂, hl₃⟩ := Finset.one_lt_card_iff.mp (Nat.one_lt_iff_ne_zero_and_ne_one.mpr ⟨hs₀, hs₁⟩) exact (eq_or_eq (hp₁ l₁ hl₁) (hp₂ l₁ hl₁) (hp₁ l₂ hl₂) (hp₂ l₂ hl₂)).resolve_right hl₃ - by_cases hs₃ : sᶜ.card = 0 + by_cases hs₃ : #sᶜ = 0 · rw [hs₃, Nat.le_zero] rw [Finset.card_compl, tsub_eq_zero_iff_le, LE.le.le_iff_eq (Finset.card_le_univ _), eq_comm, Finset.card_eq_iff_eq_univ] at hs₃ ⊢ diff --git a/Mathlib/Combinatorics/Enumerative/Catalan.lean b/Mathlib/Combinatorics/Enumerative/Catalan.lean index 0f5e376cda49f..c2ae512bb889f 100644 --- a/Mathlib/Combinatorics/Enumerative/Catalan.lean +++ b/Mathlib/Combinatorics/Enumerative/Catalan.lean @@ -186,7 +186,7 @@ theorem coe_treesOfNumNodesEq (n : ℕ) : ↑(treesOfNumNodesEq n) = { x : Tree Unit | x.numNodes = n } := Set.ext (by simp) -theorem treesOfNumNodesEq_card_eq_catalan (n : ℕ) : (treesOfNumNodesEq n).card = catalan n := by +theorem treesOfNumNodesEq_card_eq_catalan (n : ℕ) : #(treesOfNumNodesEq n) = catalan n := by induction' n using Nat.case_strong_induction_on with n ih · simp rw [treesOfNumNodesEq_succ, card_biUnion, catalan_succ'] diff --git a/Mathlib/Combinatorics/Enumerative/DoubleCounting.lean b/Mathlib/Combinatorics/Enumerative/DoubleCounting.lean index 3061124939269..9b6342a8abe9e 100644 --- a/Mathlib/Combinatorics/Enumerative/DoubleCounting.lean +++ b/Mathlib/Combinatorics/Enumerative/DoubleCounting.lean @@ -44,10 +44,10 @@ 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`. -/ -def bipartiteBelow : Finset α := s.filter fun a ↦ r a b +def bipartiteBelow : Finset α := {a ∈ s | r a b} /-- Elements of `t` which are "above" `a` according to relation `r`. -/ -def bipartiteAbove : Finset β := t.filter (r a) +def bipartiteAbove : Finset β := {b ∈ t | r a b} theorem bipartiteBelow_swap : t.bipartiteBelow (swap r) a = t.bipartiteAbove r a := rfl @@ -75,7 +75,7 @@ theorem prod_prod_bipartiteAbove_eq_prod_prod_bipartiteBelow exact prod_comm theorem sum_card_bipartiteAbove_eq_sum_card_bipartiteBelow [∀ a b, Decidable (r a b)] : - (∑ a ∈ s, (t.bipartiteAbove r a).card) = ∑ b ∈ t, (s.bipartiteBelow r b).card := by + (∑ a ∈ s, #(t.bipartiteAbove r a)) = ∑ b ∈ t, #(s.bipartiteBelow r b) := by simp_rw [card_eq_sum_ones, sum_sum_bipartiteAbove_eq_sum_sum_bipartiteBelow] section OrderedSemiring @@ -86,11 +86,11 @@ variable [OrderedSemiring R] {m n : R} Considering `r` as a bipartite graph, the LHS is a lower bound on the number of edges while the RHS is an upper bound. -/ theorem card_nsmul_le_card_nsmul [∀ a b, Decidable (r a b)] - (hm : ∀ a ∈ s, m ≤ (t.bipartiteAbove r a).card) - (hn : ∀ b ∈ t, (s.bipartiteBelow r b).card ≤ n) : s.card • m ≤ t.card • n := + (hm : ∀ a ∈ s, m ≤ #(t.bipartiteAbove r a)) + (hn : ∀ b ∈ t, #(s.bipartiteBelow r b) ≤ n) : #s • m ≤ #t • n := calc - _ ≤ ∑ a in s, ((t.bipartiteAbove r a).card : R) := s.card_nsmul_le_sum _ _ hm - _ = ∑ b in t, ((s.bipartiteBelow r b).card : R) := by + _ ≤ ∑ a in s, (#(t.bipartiteAbove r a) : R) := s.card_nsmul_le_sum _ _ hm + _ = ∑ b in t, (#(s.bipartiteBelow r b) : R) := by norm_cast; rw [sum_card_bipartiteAbove_eq_sum_card_bipartiteBelow] _ ≤ _ := t.sum_le_card_nsmul _ _ hn @@ -99,8 +99,8 @@ theorem card_nsmul_le_card_nsmul [∀ a b, Decidable (r a b)] Considering `r` as a bipartite graph, the LHS is a lower bound on the number of edges while the RHS is an upper bound. -/ theorem card_nsmul_le_card_nsmul' [∀ a b, Decidable (r a b)] - (hn : ∀ b ∈ t, n ≤ (s.bipartiteBelow r b).card) - (hm : ∀ a ∈ s, (t.bipartiteAbove r a).card ≤ m) : t.card • n ≤ s.card • m := + (hn : ∀ b ∈ t, n ≤ #(s.bipartiteBelow r b)) + (hm : ∀ a ∈ s, #(t.bipartiteAbove r a) ≤ m) : #t • n ≤ #s • m := card_nsmul_le_card_nsmul (swap r) hn hm end OrderedSemiring @@ -114,12 +114,12 @@ variable [StrictOrderedSemiring R] (r : α → β → Prop) {s : Finset α} {t : Considering `r` as a bipartite graph, the LHS is a strict lower bound on the number of edges while the RHS is an upper bound. -/ theorem card_nsmul_lt_card_nsmul_of_lt_of_le [∀ a b, Decidable (r a b)] (hs : s.Nonempty) - (hm : ∀ a ∈ s, m < (t.bipartiteAbove r a).card) - (hn : ∀ b ∈ t, (s.bipartiteBelow r b).card ≤ n) : s.card • m < t.card • n := + (hm : ∀ a ∈ s, m < #(t.bipartiteAbove r a)) + (hn : ∀ b ∈ t, #(s.bipartiteBelow r b) ≤ n) : #s • m < #t • n := calc _ = ∑ _a ∈ s, m := by rw [sum_const] - _ < ∑ a ∈ s, ((t.bipartiteAbove r a).card : R) := sum_lt_sum_of_nonempty hs hm - _ = ∑ b in t, ((s.bipartiteBelow r b).card : R) := by + _ < ∑ a ∈ s, (#(t.bipartiteAbove r a) : R) := sum_lt_sum_of_nonempty hs hm + _ = ∑ b in t, (#(s.bipartiteBelow r b) : R) := by norm_cast; rw [sum_card_bipartiteAbove_eq_sum_card_bipartiteBelow] _ ≤ _ := t.sum_le_card_nsmul _ _ hn @@ -128,11 +128,11 @@ theorem card_nsmul_lt_card_nsmul_of_lt_of_le [∀ a b, Decidable (r a b)] (hs : Considering `r` as a bipartite graph, the LHS is a lower bound on the number of edges while the RHS is a strict upper bound. -/ theorem card_nsmul_lt_card_nsmul_of_le_of_lt [∀ a b, Decidable (r a b)] (ht : t.Nonempty) - (hm : ∀ a ∈ s, m ≤ (t.bipartiteAbove r a).card) - (hn : ∀ b ∈ t, (s.bipartiteBelow r b).card < n) : s.card • m < t.card • n := + (hm : ∀ a ∈ s, m ≤ #(t.bipartiteAbove r a)) + (hn : ∀ b ∈ t, #(s.bipartiteBelow r b) < n) : #s • m < #t • n := calc - _ ≤ ∑ a in s, ((t.bipartiteAbove r a).card : R) := s.card_nsmul_le_sum _ _ hm - _ = ∑ b in t, ((s.bipartiteBelow r b).card : R) := by + _ ≤ ∑ a in s, (#(t.bipartiteAbove r a) : R) := s.card_nsmul_le_sum _ _ hm + _ = ∑ b in t, (#(s.bipartiteBelow r b) : R) := by norm_cast; rw [sum_card_bipartiteAbove_eq_sum_card_bipartiteBelow] _ < ∑ _b ∈ t, n := sum_lt_sum_of_nonempty ht hn _ = _ := sum_const _ @@ -142,8 +142,8 @@ theorem card_nsmul_lt_card_nsmul_of_le_of_lt [∀ a b, Decidable (r a b)] (ht : Considering `r` as a bipartite graph, the LHS is a strict lower bound on the number of edges while the RHS is an upper bound. -/ theorem card_nsmul_lt_card_nsmul_of_lt_of_le' [∀ a b, Decidable (r a b)] (ht : t.Nonempty) - (hn : ∀ b ∈ t, n < (s.bipartiteBelow r b).card) - (hm : ∀ a ∈ s, (t.bipartiteAbove r a).card ≤ m) : t.card • n < s.card • m := + (hn : ∀ b ∈ t, n < #(s.bipartiteBelow r b)) + (hm : ∀ a ∈ s, #(t.bipartiteAbove r a) ≤ m) : #t • n < #s • m := card_nsmul_lt_card_nsmul_of_lt_of_le (swap r) ht hn hm /-- **Double counting** argument. @@ -151,8 +151,8 @@ theorem card_nsmul_lt_card_nsmul_of_lt_of_le' [∀ a b, Decidable (r a b)] (ht : Considering `r` as a bipartite graph, the LHS is a lower bound on the number of edges while the RHS is a strict upper bound. -/ theorem card_nsmul_lt_card_nsmul_of_le_of_lt' [∀ a b, Decidable (r a b)] (hs : s.Nonempty) - (hn : ∀ b ∈ t, n ≤ (s.bipartiteBelow r b).card) - (hm : ∀ a ∈ s, (t.bipartiteAbove r a).card < m) : t.card • n < s.card • m := + (hn : ∀ b ∈ t, n ≤ #(s.bipartiteBelow r b)) + (hm : ∀ a ∈ s, #(t.bipartiteAbove r a) < m) : #t • n < #s • m := card_nsmul_lt_card_nsmul_of_le_of_lt (swap r) hs hn hm end StrictOrderedSemiring @@ -162,25 +162,25 @@ end StrictOrderedSemiring Considering `r` as a bipartite graph, the LHS is a lower bound on the number of edges while the RHS is an upper bound. -/ theorem card_mul_le_card_mul [∀ a b, Decidable (r a b)] - (hm : ∀ a ∈ s, m ≤ (t.bipartiteAbove r a).card) - (hn : ∀ b ∈ t, (s.bipartiteBelow r b).card ≤ n) : s.card * m ≤ t.card * n := + (hm : ∀ a ∈ s, m ≤ #(t.bipartiteAbove r a)) + (hn : ∀ b ∈ t, #(s.bipartiteBelow r b) ≤ n) : #s * m ≤ #t * n := card_nsmul_le_card_nsmul _ hm hn theorem card_mul_le_card_mul' [∀ a b, Decidable (r a b)] - (hn : ∀ b ∈ t, n ≤ (s.bipartiteBelow r b).card) - (hm : ∀ a ∈ s, (t.bipartiteAbove r a).card ≤ m) : t.card * n ≤ s.card * m := + (hn : ∀ b ∈ t, n ≤ #(s.bipartiteBelow r b)) + (hm : ∀ a ∈ s, #(t.bipartiteAbove r a) ≤ m) : #t * n ≤ #s * m := card_nsmul_le_card_nsmul' _ hn hm theorem card_mul_eq_card_mul [∀ a b, Decidable (r a b)] - (hm : ∀ a ∈ s, (t.bipartiteAbove r a).card = m) - (hn : ∀ b ∈ t, (s.bipartiteBelow r b).card = n) : s.card * m = t.card * n := + (hm : ∀ a ∈ s, #(t.bipartiteAbove r a) = m) + (hn : ∀ b ∈ t, #(s.bipartiteBelow r b) = n) : #s * m = #t * n := (card_mul_le_card_mul _ (fun a ha ↦ (hm a ha).ge) fun b hb ↦ (hn b hb).le).antisymm <| card_mul_le_card_mul' _ (fun a ha ↦ (hn a ha).ge) fun b hb ↦ (hm b hb).le theorem card_le_card_of_forall_subsingleton (hs : ∀ a ∈ s, ∃ b, b ∈ t ∧ r a b) - (ht : ∀ b ∈ t, ({ a ∈ s | r a b } : Set α).Subsingleton) : s.card ≤ t.card := by + (ht : ∀ b ∈ t, ({ a ∈ s | r a b } : Set α).Subsingleton) : #s ≤ #t := by classical - rw [← mul_one s.card, ← mul_one t.card] + rw [← mul_one #s, ← mul_one #t] exact card_mul_le_card_mul r (fun a h ↦ card_pos.2 (by rw [← coe_nonempty, coe_bipartiteAbove] @@ -190,7 +190,7 @@ theorem card_le_card_of_forall_subsingleton (hs : ∀ a ∈ s, ∃ b, b ∈ t exact ht _ h)) theorem card_le_card_of_forall_subsingleton' (ht : ∀ b ∈ t, ∃ a, a ∈ s ∧ r a b) - (hs : ∀ a ∈ s, ({ b ∈ t | r a b } : Set β).Subsingleton) : t.card ≤ s.card := + (hs : ∀ a ∈ s, ({ b ∈ t | r a b } : Set β).Subsingleton) : #t ≤ #s := card_le_card_of_forall_subsingleton (swap r) ht hs end Bipartite diff --git a/Mathlib/Combinatorics/Hall/Basic.lean b/Mathlib/Combinatorics/Hall/Basic.lean index 0d0b9a82657a9..da15e7bee4a35 100644 --- a/Mathlib/Combinatorics/Hall/Basic.lean +++ b/Mathlib/Combinatorics/Hall/Basic.lean @@ -50,8 +50,7 @@ The core of this module is constructing the inverse system: for every finite sub Hall's Marriage Theorem, indexed families -/ - -open Finset CategoryTheory +open Finset Function CategoryTheory universe u v @@ -71,7 +70,7 @@ def hallMatchingsOn.restrict {ι : Type u} {α : Type v} (t : ι → Finset α) /-- When the Hall condition is satisfied, the set of matchings on a finite set is nonempty. This is where `Finset.all_card_le_biUnion_card_iff_existsInjective'` comes into the argument. -/ theorem hallMatchingsOn.nonempty {ι : Type u} {α : Type v} [DecidableEq α] (t : ι → Finset α) - (h : ∀ s : Finset ι, s.card ≤ (s.biUnion t).card) (ι' : Finset ι) : + (h : ∀ s : Finset ι, #s ≤ #(s.biUnion t)) (ι' : Finset ι) : Nonempty (hallMatchingsOn t ι') := by classical refine ⟨Classical.indefiniteDescription _ ?_⟩ @@ -115,7 +114,7 @@ which has the additional constraint that `ι` is a `Fintype`. -/ theorem Finset.all_card_le_biUnion_card_iff_exists_injective {ι : Type u} {α : Type v} [DecidableEq α] (t : ι → Finset α) : - (∀ s : Finset ι, s.card ≤ (s.biUnion t).card) ↔ + (∀ s : Finset ι, #s ≤ #(s.biUnion t)) ↔ ∃ f : ι → α, Function.Injective f ∧ ∀ x, f x ∈ t x := by constructor · intro h @@ -178,10 +177,10 @@ Note: if `[Fintype β]`, then there exist instances for `[∀ (a : α), Fintype -/ theorem Fintype.all_card_le_rel_image_card_iff_exists_injective {α : Type u} {β : Type v} [DecidableEq β] (r : α → β → Prop) [∀ a : α, Fintype (Rel.image r {a})] : - (∀ A : Finset α, A.card ≤ Fintype.card (Rel.image r A)) ↔ + (∀ A : Finset α, #A ≤ Fintype.card (Rel.image r A)) ↔ ∃ f : α → β, Function.Injective f ∧ ∀ x, r x (f x) := by let r' a := (Rel.image r {a}).toFinset - have h : ∀ A : Finset α, Fintype.card (Rel.image r A) = (A.biUnion r').card := by + have h : ∀ A : Finset α, Fintype.card (Rel.image r A) = #(A.biUnion r') := by intro A rw [← Set.toFinset_card] apply congr_arg @@ -203,11 +202,10 @@ rather than `Rel.image`. -/ theorem Fintype.all_card_le_filter_rel_iff_exists_injective {α : Type u} {β : Type v} [Fintype β] (r : α → β → Prop) [∀ a, DecidablePred (r a)] : - (∀ A : Finset α, A.card ≤ (univ.filter fun b : β => ∃ a ∈ A, r a b).card) ↔ - ∃ f : α → β, Function.Injective f ∧ ∀ x, r x (f x) := by + (∀ A : Finset α, #A ≤ #{b | ∃ a ∈ A, r a b}) ↔ ∃ f : α → β, Injective f ∧ ∀ x, r x (f x) := by haveI := Classical.decEq β - let r' a := univ.filter fun b => r a b - have h : ∀ A : Finset α, (univ.filter fun b : β => ∃ a ∈ A, r a b) = A.biUnion r' := by + let r' a : Finset β := {b | r a b} + have h : ∀ A : Finset α, ({b | ∃ a ∈ A, r a b} : Finset _) = A.biUnion r' := by intro A ext b simp [r'] diff --git a/Mathlib/Combinatorics/Hall/Finite.lean b/Mathlib/Combinatorics/Hall/Finite.lean index aadc2cf7010b8..a81d2cf78e3d0 100644 --- a/Mathlib/Combinatorics/Hall/Finite.lean +++ b/Mathlib/Combinatorics/Hall/Finite.lean @@ -46,13 +46,13 @@ section Fintype variable [Fintype ι] theorem hall_cond_of_erase {x : ι} (a : α) - (ha : ∀ s : Finset ι, s.Nonempty → s ≠ univ → s.card < (s.biUnion t).card) - (s' : Finset { x' : ι | x' ≠ x }) : s'.card ≤ (s'.biUnion fun x' => (t x').erase a).card := by + (ha : ∀ s : Finset ι, s.Nonempty → s ≠ univ → #s < #(s.biUnion t)) + (s' : Finset { x' : ι | x' ≠ x }) : #s' ≤ #(s'.biUnion fun x' => (t x').erase a) := by haveI := Classical.decEq ι specialize ha (s'.image fun z => z.1) rw [image_nonempty, Finset.card_image_of_injective s' Subtype.coe_injective] at ha by_cases he : s'.Nonempty - · have ha' : s'.card < (s'.biUnion fun x => t x).card := by + · have ha' : #s' < #(s'.biUnion fun x => t x) := by convert ha he fun h => by simpa [← h] using mem_univ x using 2 ext x simp only [mem_image, mem_biUnion, exists_prop, SetCoe.exists, exists_and_right, @@ -68,18 +68,18 @@ theorem hall_cond_of_erase {x : ι} (a : α) simp /-- First case of the inductive step: assuming that -`∀ (s : Finset ι), s.Nonempty → s ≠ univ → s.card < (s.biUnion t).card` +`∀ (s : Finset ι), s.Nonempty → s ≠ univ → #s < #(s.biUnion t)` and that the statement of **Hall's Marriage Theorem** is true for all `ι'` of cardinality ≤ `n`, then it is true for `ι` of cardinality `n + 1`. -/ theorem hall_hard_inductive_step_A {n : ℕ} (hn : Fintype.card ι = n + 1) - (ht : ∀ s : Finset ι, s.card ≤ (s.biUnion t).card) + (ht : ∀ s : Finset ι, #s ≤ #(s.biUnion t)) (ih : ∀ {ι' : Type u} [Fintype ι'] (t' : ι' → Finset α), Fintype.card ι' ≤ n → - (∀ s' : Finset ι', s'.card ≤ (s'.biUnion t').card) → + (∀ s' : Finset ι', #s' ≤ #(s'.biUnion t')) → ∃ f : ι' → α, Function.Injective f ∧ ∀ x, f x ∈ t' x) - (ha : ∀ s : Finset ι, s.Nonempty → s ≠ univ → s.card < (s.biUnion t).card) : + (ha : ∀ s : Finset ι, s.Nonempty → s ≠ univ → #s < #(s.biUnion t)) : ∃ f : ι → α, Function.Injective f ∧ ∀ x, f x ∈ t x := by haveI : Nonempty ι := Fintype.card_pos_iff.mp (hn.symm ▸ Nat.succ_pos _) haveI := Classical.decEq ι @@ -89,7 +89,7 @@ theorem hall_hard_inductive_step_A {n : ℕ} (hn : Fintype.card ι = n + 1) rw [← Finset.card_pos] calc 0 < 1 := Nat.one_pos - _ ≤ (Finset.biUnion {x} t).card := ht {x} + _ ≤ #(.biUnion {x} t) := ht {x} _ = (t x).card := by rw [Finset.singleton_biUnion] choose y hy using tx_ne @@ -118,8 +118,8 @@ theorem hall_hard_inductive_step_A {n : ℕ} (hn : Fintype.card ι = n + 1) exact hfr.2 theorem hall_cond_of_restrict {ι : Type u} {t : ι → Finset α} {s : Finset ι} - (ht : ∀ s : Finset ι, s.card ≤ (s.biUnion t).card) (s' : Finset (s : Set ι)) : - s'.card ≤ (s'.biUnion fun a' => t a').card := by + (ht : ∀ s : Finset ι, #s ≤ #(s.biUnion t)) (s' : Finset (s : Set ι)) : + #s' ≤ #(s'.biUnion fun a' => t a') := by classical rw [← card_image_of_injective s' Subtype.coe_injective] convert ht (s'.image fun z => z.1) using 1 @@ -128,15 +128,15 @@ theorem hall_cond_of_restrict {ι : Type u} {t : ι → Finset α} {s : Finset simp theorem hall_cond_of_compl {ι : Type u} {t : ι → Finset α} {s : Finset ι} - (hus : s.card = (s.biUnion t).card) (ht : ∀ s : Finset ι, s.card ≤ (s.biUnion t).card) - (s' : Finset (sᶜ : Set ι)) : s'.card ≤ (s'.biUnion fun x' => t x' \ s.biUnion t).card := by + (hus : #s = #(s.biUnion t)) (ht : ∀ s : Finset ι, #s ≤ #(s.biUnion t)) + (s' : Finset (sᶜ : Set ι)) : #s' ≤ #(s'.biUnion fun x' => t x' \ s.biUnion t) := by haveI := Classical.decEq ι have disj : Disjoint s (s'.image fun z => z.1) := by simp only [disjoint_left, not_exists, mem_image, exists_prop, SetCoe.exists, exists_and_right, exists_eq_right, Subtype.coe_mk] intro x hx hc _ exact absurd hx hc - have : s'.card = (s ∪ s'.image fun z => z.1).card - s.card := by + have : #s' = #(s ∪ s'.image fun z => z.1) - #s := by simp [disj, card_image_of_injective _ Subtype.coe_injective, Nat.add_sub_cancel_left] rw [this, hus] refine (Nat.sub_le_sub_right (ht _) _).trans ?_ @@ -152,18 +152,18 @@ theorem hall_cond_of_compl {ι : Type u} {t : ι → Finset α} {s : Finset ι} apply subset_union_left /-- Second case of the inductive step: assuming that -`∃ (s : Finset ι), s ≠ univ → s.card = (s.biUnion t).card` +`∃ (s : Finset ι), s ≠ univ → #s = #(s.biUnion t)` and that the statement of **Hall's Marriage Theorem** is true for all `ι'` of cardinality ≤ `n`, then it is true for `ι` of cardinality `n + 1`. -/ theorem hall_hard_inductive_step_B {n : ℕ} (hn : Fintype.card ι = n + 1) - (ht : ∀ s : Finset ι, s.card ≤ (s.biUnion t).card) + (ht : ∀ s : Finset ι, #s ≤ #(s.biUnion t)) (ih : ∀ {ι' : Type u} [Fintype ι'] (t' : ι' → Finset α), Fintype.card ι' ≤ n → - (∀ s' : Finset ι', s'.card ≤ (s'.biUnion t').card) → + (∀ s' : Finset ι', #s' ≤ #(s'.biUnion t')) → ∃ f : ι' → α, Function.Injective f ∧ ∀ x, f x ∈ t' x) - (s : Finset ι) (hs : s.Nonempty) (hns : s ≠ univ) (hus : s.card = (s.biUnion t).card) : + (s : Finset ι) (hs : s.Nonempty) (hns : s ≠ univ) (hus : #s = #(s.biUnion t)) : ∃ f : ι → α, Function.Injective f ∧ ∀ x, f x ∈ t x := by haveI := Classical.decEq ι -- Restrict to `s` @@ -171,7 +171,7 @@ theorem hall_hard_inductive_step_B {n : ℕ} (hn : Fintype.card ι = n + 1) have card_ι'_le : Fintype.card s ≤ n := by apply Nat.le_of_lt_succ calc - Fintype.card s = s.card := Fintype.card_coe _ + Fintype.card s = #s := Fintype.card_coe _ _ < Fintype.card ι := (card_lt_iff_ne_univ _).mpr hns _ = n.succ := hn let t' : s → Finset α := fun x' => t x' @@ -214,7 +214,7 @@ variable [Finite ι] /-- Here we combine the two inductive steps into a full strong induction proof, completing the proof the harder direction of **Hall's Marriage Theorem**. -/ -theorem hall_hard_inductive (ht : ∀ s : Finset ι, s.card ≤ (s.biUnion t).card) : +theorem hall_hard_inductive (ht : ∀ s : Finset ι, #s ≤ #(s.biUnion t)) : ∃ f : ι → α, Function.Injective f ∧ ∀ x, f x ∈ t x := by cases nonempty_fintype ι induction' hn : Fintype.card ι using Nat.strong_induction_on with n ih generalizing ι @@ -222,11 +222,11 @@ theorem hall_hard_inductive (ht : ∀ s : Finset ι, s.card ≤ (s.biUnion t).ca · rw [Fintype.card_eq_zero_iff] at hn exact ⟨isEmptyElim, isEmptyElim, isEmptyElim⟩ · have ih' : ∀ (ι' : Type u) [Fintype ι'] (t' : ι' → Finset α), Fintype.card ι' ≤ n → - (∀ s' : Finset ι', s'.card ≤ (s'.biUnion t').card) → + (∀ s' : Finset ι', #s' ≤ #(s'.biUnion t')) → ∃ f : ι' → α, Function.Injective f ∧ ∀ x, f x ∈ t' x := by intro ι' _ _ hι' ht' exact ih _ (Nat.lt_succ_of_le hι') ht' _ rfl - by_cases h : ∀ s : Finset ι, s.Nonempty → s ≠ univ → s.card < (s.biUnion t).card + by_cases h : ∀ s : Finset ι, s.Nonempty → s ≠ univ → #s < #(s.biUnion t) · refine hall_hard_inductive_step_A hn ht (@fun ι' => ih' ι') h · push_neg at h rcases h with ⟨s, sne, snu, sle⟩ @@ -245,7 +245,7 @@ where the `Finite ι` constraint is removed. -/ theorem Finset.all_card_le_biUnion_card_iff_existsInjective' {ι α : Type*} [Finite ι] [DecidableEq α] (t : ι → Finset α) : - (∀ s : Finset ι, s.card ≤ (s.biUnion t).card) ↔ + (∀ s : Finset ι, #s ≤ #(s.biUnion t)) ↔ ∃ f : ι → α, Function.Injective f ∧ ∀ x, f x ∈ t x := by constructor · exact HallMarriageTheorem.hall_hard_inductive diff --git a/Mathlib/Combinatorics/Pigeonhole.lean b/Mathlib/Combinatorics/Pigeonhole.lean index b91d6c658fd57..fe51886fb9187 100644 --- a/Mathlib/Combinatorics/Pigeonhole.lean +++ b/Mathlib/Combinatorics/Pigeonhole.lean @@ -32,14 +32,14 @@ The versions vary by: * using a function between `Fintype`s or a function between possibly infinite types restricted to `Finset`s; -* counting pigeons by a general weight function (`∑ x ∈ s, w x`) or by heads (`Finset.card s`); +* counting pigeons by a general weight function (`∑ x ∈ s, w x`) or by heads (`#s`); * using strict or non-strict inequalities; * establishing upper or lower estimate on the number (or the total weight) of the pigeons in one pigeonhole; * in case when we count pigeons by some weight function `w` and consider a function `f` between `Finset`s `s` and `t`, we can either assume that each pigeon is in one of the pigeonholes (`∀ x ∈ s, f x ∈ t`), or assume that for `y ∉ t`, the total weight of the pigeons in this - pigeonhole `∑ x ∈ s.filter (fun x ↦ f x = y), w x` is nonpositive or nonnegative depending on + pigeonhole `∑ x ∈ s with f x = y, w x` is nonpositive or nonnegative depending on the inequality we are proving. Lemma names follow `mathlib` convention (e.g., @@ -80,15 +80,14 @@ variations of this theorem. The principle is formalized in the following way, see `Finset.exists_lt_sum_fiber_of_maps_to_of_nsmul_lt_sum`: if `f : α → β` is a function which maps all -elements of `s : Finset α` to `t : Finset β` and `card t • b < ∑ x ∈ s, w x`, where `w : α → M` is +elements of `s : Finset α` to `t : Finset β` and `#t • b < ∑ x ∈ s, w x`, where `w : α → M` is a weight function taking values in a `LinearOrderedCancelAddCommMonoid`, then for some `y ∈ t`, the sum of the weights of all `x ∈ s` such that `f x = y` is greater than `b`. There are a few bits we can change in this theorem: * reverse all inequalities, with obvious adjustments to the name; -* replace the assumption `∀ a ∈ s, f a ∈ t` with - `∀ y ∉ t, (∑ x ∈ s.filter (fun x ↦ f x = y), w x) ≤ 0`, +* replace the assumption `∀ a ∈ s, f a ∈ t` with `∀ y ∉ t, ∑ x ∈ s with f x = y, w x ≤ 0`, and replace `of_maps_to` with `of_sum_fiber_nonpos` in the name; * use non-strict inequalities assuming `t` is nonempty. @@ -109,7 +108,7 @@ if the total weight of a finite set of pigeons is greater than `n • b`, and th `n` pigeonholes, then for some pigeonhole, the total weight of the pigeons in this pigeonhole is greater than `b`. -/ theorem exists_lt_sum_fiber_of_maps_to_of_nsmul_lt_sum (hf : ∀ a ∈ s, f a ∈ t) - (hb : t.card • b < ∑ x ∈ s, w x) : ∃ y ∈ t, b < ∑ x ∈ s.filter fun x => f x = y, w x := + (hb : #t • b < ∑ x ∈ s, w x) : ∃ y ∈ t, b < ∑ x ∈ s with f x = y, w x := exists_lt_of_sum_lt <| by simpa only [sum_fiberwise_of_maps_to hf, sum_const] /-- The pigeonhole principle for finitely many pigeons counted by weight, strict inequality version: @@ -117,7 +116,7 @@ if the total weight of a finite set of pigeons is less than `n • b`, and they pigeonholes, then for some pigeonhole, the total weight of the pigeons in this pigeonhole is less than `b`. -/ theorem exists_sum_fiber_lt_of_maps_to_of_sum_lt_nsmul (hf : ∀ a ∈ s, f a ∈ t) - (hb : ∑ x ∈ s, w x < t.card • b) : ∃ y ∈ t, ∑ x ∈ s.filter fun x => f x = y, w x < b := + (hb : ∑ x ∈ s, w x < #t • b) : ∃ y ∈ t, ∑ x ∈ s with f x = y, w x < b := exists_lt_sum_fiber_of_maps_to_of_nsmul_lt_sum (M := Mᵒᵈ) hf hb /-- The pigeonhole principle for finitely many pigeons counted by weight, strict inequality version: @@ -126,13 +125,12 @@ pigeonholes, and for all but `n` pigeonholes the total weight of the pigeons the then for at least one of these `n` pigeonholes, the total weight of the pigeons in this pigeonhole is greater than `b`. -/ theorem exists_lt_sum_fiber_of_sum_fiber_nonpos_of_nsmul_lt_sum - (ht : ∀ y ∉ t, ∑ x ∈ s.filter fun x => f x = y, w x ≤ 0) - (hb : t.card • b < ∑ x ∈ s, w x) : ∃ y ∈ t, b < ∑ x ∈ s.filter fun x => f x = y, w x := + (ht : ∀ y ∉ t, ∑ x ∈ s with f x = y, w x ≤ 0) + (hb : #t • b < ∑ x ∈ s, w x) : ∃ y ∈ t, b < ∑ x ∈ s with f x = y, w x := exists_lt_of_sum_lt <| calc ∑ _y ∈ t, b < ∑ x ∈ s, w x := by simpa - _ ≤ ∑ y ∈ t, ∑ x ∈ s.filter fun x => f x = y, w x := - sum_le_sum_fiberwise_of_sum_fiber_nonpos ht + _ ≤ ∑ y ∈ t, ∑ x ∈ s with f x = y, w x := sum_le_sum_fiberwise_of_sum_fiber_nonpos ht /-- The pigeonhole principle for finitely many pigeons counted by weight, strict inequality version: if the total weight of a finite set of pigeons is less than `n • b`, they are sorted into some @@ -140,8 +138,8 @@ pigeonholes, and for all but `n` pigeonholes the total weight of the pigeons the then for at least one of these `n` pigeonholes, the total weight of the pigeons in this pigeonhole is less than `b`. -/ theorem exists_sum_fiber_lt_of_sum_fiber_nonneg_of_sum_lt_nsmul - (ht : ∀ y ∉ t, (0 : M) ≤ ∑ x ∈ s.filter fun x => f x = y, w x) - (hb : ∑ x ∈ s, w x < t.card • b) : ∃ y ∈ t, ∑ x ∈ s.filter fun x => f x = y, w x < b := + (ht : ∀ y ∉ t, (0 : M) ≤ ∑ x ∈ s with f x = y, w x) (hb : ∑ x ∈ s, w x < #t • b) : + ∃ y ∈ t, ∑ x ∈ s with f x = y, w x < b := exists_lt_sum_fiber_of_sum_fiber_nonpos_of_nsmul_lt_sum (M := Mᵒᵈ) ht hb /-! @@ -154,7 +152,7 @@ version: if the total weight of a finite set of pigeons is greater than or equal they are sorted into `n > 0` pigeonholes, then for some pigeonhole, the total weight of the pigeons in this pigeonhole is greater than or equal to `b`. -/ theorem exists_le_sum_fiber_of_maps_to_of_nsmul_le_sum (hf : ∀ a ∈ s, f a ∈ t) (ht : t.Nonempty) - (hb : t.card • b ≤ ∑ x ∈ s, w x) : ∃ y ∈ t, b ≤ ∑ x ∈ s.filter fun x => f x = y, w x := + (hb : #t • b ≤ ∑ x ∈ s, w x) : ∃ y ∈ t, b ≤ ∑ x ∈ s with f x = y, w x := exists_le_of_sum_le ht <| by simpa only [sum_fiberwise_of_maps_to hf, sum_const] /-- The pigeonhole principle for finitely many pigeons counted by weight, non-strict inequality @@ -162,7 +160,7 @@ version: if the total weight of a finite set of pigeons is less than or equal to are sorted into `n > 0` pigeonholes, then for some pigeonhole, the total weight of the pigeons in this pigeonhole is less than or equal to `b`. -/ theorem exists_sum_fiber_le_of_maps_to_of_sum_le_nsmul (hf : ∀ a ∈ s, f a ∈ t) (ht : t.Nonempty) - (hb : ∑ x ∈ s, w x ≤ t.card • b) : ∃ y ∈ t, ∑ x ∈ s.filter fun x => f x = y, w x ≤ b := + (hb : ∑ x ∈ s, w x ≤ #t • b) : ∃ y ∈ t, ∑ x ∈ s with f x = y, w x ≤ b := exists_le_sum_fiber_of_maps_to_of_nsmul_le_sum (M := Mᵒᵈ) hf ht hb /-- The pigeonhole principle for finitely many pigeons counted by weight, non-strict inequality @@ -171,12 +169,12 @@ are sorted into some pigeonholes, and for all but `n > 0` pigeonholes the total pigeons there is nonpositive, then for at least one of these `n` pigeonholes, the total weight of the pigeons in this pigeonhole is greater than or equal to `b`. -/ theorem exists_le_sum_fiber_of_sum_fiber_nonpos_of_nsmul_le_sum - (hf : ∀ y ∉ t, ∑ x ∈ s.filter fun x => f x = y, w x ≤ 0) (ht : t.Nonempty) - (hb : t.card • b ≤ ∑ x ∈ s, w x) : ∃ y ∈ t, b ≤ ∑ x ∈ s.filter fun x => f x = y, w x := + (hf : ∀ y ∉ t, ∑ x ∈ s with f x = y, w x ≤ 0) (ht : t.Nonempty) + (hb : #t • b ≤ ∑ x ∈ s, w x) : ∃ y ∈ t, b ≤ ∑ x ∈ s with f x = y, w x := exists_le_of_sum_le ht <| calc ∑ _y ∈ t, b ≤ ∑ x ∈ s, w x := by simpa - _ ≤ ∑ y ∈ t, ∑ x ∈ s.filter fun x => f x = y, w x := + _ ≤ ∑ y ∈ t, ∑ x ∈ s with f x = y, w x := sum_le_sum_fiberwise_of_sum_fiber_nonpos hf /-- The pigeonhole principle for finitely many pigeons counted by weight, non-strict inequality @@ -185,8 +183,8 @@ sorted into some pigeonholes, and for all but `n > 0` pigeonholes the total weig there is nonnegative, then for at least one of these `n` pigeonholes, the total weight of the pigeons in this pigeonhole is less than or equal to `b`. -/ theorem exists_sum_fiber_le_of_sum_fiber_nonneg_of_sum_le_nsmul - (hf : ∀ y ∉ t, (0 : M) ≤ ∑ x ∈ s.filter fun x => f x = y, w x) (ht : t.Nonempty) - (hb : ∑ x ∈ s, w x ≤ t.card • b) : ∃ y ∈ t, ∑ x ∈ s.filter fun x => f x = y, w x ≤ b := + (hf : ∀ y ∉ t, (0 : M) ≤ ∑ x ∈ s with f x = y, w x) (ht : t.Nonempty) + (hb : ∑ x ∈ s, w x ≤ #t • b) : ∃ y ∈ t, ∑ x ∈ s with f x = y, w x ≤ b := exists_le_sum_fiber_of_sum_fiber_nonpos_of_nsmul_le_sum (M := Mᵒᵈ) hf ht hb end @@ -214,7 +212,7 @@ So, we prove four theorems: `Finset.exists_lt_card_fiber_of_maps_to_of_mul_lt_ca /-- The pigeonhole principle for finitely many pigeons counted by heads: there is a pigeonhole with at least as many pigeons as the ceiling of the average number of pigeons across all pigeonholes. -/ theorem exists_lt_card_fiber_of_nsmul_lt_card_of_maps_to (hf : ∀ a ∈ s, f a ∈ t) - (ht : t.card • b < s.card) : ∃ y ∈ t, b < (s.filter fun x => f x = y).card := by + (ht : #t • b < #s) : ∃ y ∈ t, b < #{x ∈ s | f x = y} := by simp_rw [cast_card] at ht ⊢ exact exists_lt_sum_fiber_of_maps_to_of_nsmul_lt_sum hf ht @@ -223,16 +221,16 @@ at least as many pigeons as the ceiling of the average number of pigeons across ("The maximum is at least the mean" specialized to integers.) More formally, given a function between finite sets `s` and `t` and a natural number `n` such that -`card t * n < card s`, there exists `y ∈ t` such that its preimage in `s` has more than `n` +`#t * n < #s`, there exists `y ∈ t` such that its preimage in `s` has more than `n` elements. -/ theorem exists_lt_card_fiber_of_mul_lt_card_of_maps_to (hf : ∀ a ∈ s, f a ∈ t) - (hn : t.card * n < s.card) : ∃ y ∈ t, n < (s.filter fun x => f x = y).card := + (hn : #t * n < #s) : ∃ y ∈ t, n < #{x ∈ s | f x = y} := exists_lt_card_fiber_of_nsmul_lt_card_of_maps_to hf hn /-- The pigeonhole principle for finitely many pigeons counted by heads: there is a pigeonhole with at most as many pigeons as the floor of the average number of pigeons across all pigeonholes. -/ -theorem exists_card_fiber_lt_of_card_lt_nsmul (ht : ↑s.card < t.card • b) : - ∃ y ∈ t, ↑(s.filter fun x => f x = y).card < b := by +theorem exists_card_fiber_lt_of_card_lt_nsmul (ht : #s < #t • b) : + ∃ y ∈ t, #{x ∈ s | f x = y} < b := by simp_rw [cast_card] at ht ⊢ exact exists_sum_fiber_lt_of_sum_fiber_nonneg_of_sum_lt_nsmul @@ -243,35 +241,34 @@ at most as many pigeons as the floor of the average number of pigeons across all minimum is at most the mean" specialized to integers.) More formally, given a function `f`, a finite sets `s` in its domain, a finite set `t` in its -codomain, and a natural number `n` such that `card s < card t * n`, there exists `y ∈ t` such that +codomain, and a natural number `n` such that `#s < #t * n`, there exists `y ∈ t` such that its preimage in `s` has less than `n` elements. -/ -theorem exists_card_fiber_lt_of_card_lt_mul (hn : s.card < t.card * n) : - ∃ y ∈ t, (s.filter fun x => f x = y).card < n := +theorem exists_card_fiber_lt_of_card_lt_mul (hn : #s < #t * n) : ∃ y ∈ t, #{x ∈ s | f x = y} < n := exists_card_fiber_lt_of_card_lt_nsmul hn /-- The pigeonhole principle for finitely many pigeons counted by heads: given a function between -finite sets `s` and `t` and a number `b` such that `card t • b ≤ card s`, there exists `y ∈ t` such +finite sets `s` and `t` and a number `b` such that `#t • b ≤ #s`, there exists `y ∈ t` such that its preimage in `s` has at least `b` elements. See also `Finset.exists_lt_card_fiber_of_nsmul_lt_card_of_maps_to` for a stronger statement. -/ theorem exists_le_card_fiber_of_nsmul_le_card_of_maps_to (hf : ∀ a ∈ s, f a ∈ t) (ht : t.Nonempty) - (hb : t.card • b ≤ s.card) : ∃ y ∈ t, b ≤ (s.filter fun x => f x = y).card := by + (hb : #t • b ≤ #s) : ∃ y ∈ t, b ≤ #{x ∈ s | f x = y} := by simp_rw [cast_card] at hb ⊢ exact exists_le_sum_fiber_of_maps_to_of_nsmul_le_sum hf ht hb /-- The pigeonhole principle for finitely many pigeons counted by heads: given a function between -finite sets `s` and `t` and a natural number `b` such that `card t * n ≤ card s`, there exists +finite sets `s` and `t` and a natural number `b` such that `#t * n ≤ #s`, there exists `y ∈ t` such that its preimage in `s` has at least `n` elements. See also `Finset.exists_lt_card_fiber_of_mul_lt_card_of_maps_to` for a stronger statement. -/ theorem exists_le_card_fiber_of_mul_le_card_of_maps_to (hf : ∀ a ∈ s, f a ∈ t) (ht : t.Nonempty) - (hn : t.card * n ≤ s.card) : ∃ y ∈ t, n ≤ (s.filter fun x => f x = y).card := + (hn : #t * n ≤ #s) : ∃ y ∈ t, n ≤ #{x ∈ s | f x = y} := exists_le_card_fiber_of_nsmul_le_card_of_maps_to hf ht hn /-- The pigeonhole principle for finitely many pigeons counted by heads: given a function `f`, a -finite sets `s` and `t`, and a number `b` such that `card s ≤ card t • b`, there exists `y ∈ t` such +finite sets `s` and `t`, and a number `b` such that `#s ≤ #t • b`, there exists `y ∈ t` such that its preimage in `s` has no more than `b` elements. See also `Finset.exists_card_fiber_lt_of_card_lt_nsmul` for a stronger statement. -/ -theorem exists_card_fiber_le_of_card_le_nsmul (ht : t.Nonempty) (hb : ↑s.card ≤ t.card • b) : - ∃ y ∈ t, ↑(s.filter fun x => f x = y).card ≤ b := by +theorem exists_card_fiber_le_of_card_le_nsmul (ht : t.Nonempty) (hb : #s ≤ #t • b) : + ∃ y ∈ t, #{x ∈ s | f x = y} ≤ b := by simp_rw [cast_card] at hb ⊢ refine exists_sum_fiber_le_of_sum_fiber_nonneg_of_sum_le_nsmul @@ -279,10 +276,10 @@ theorem exists_card_fiber_le_of_card_le_nsmul (ht : t.Nonempty) (hb : ↑s.card /-- The pigeonhole principle for finitely many pigeons counted by heads: given a function `f`, a finite sets `s` in its domain, a finite set `t` in its codomain, and a natural number `n` such that -`card s ≤ card t * n`, there exists `y ∈ t` such that its preimage in `s` has no more than `n` +`#s ≤ #t * n`, there exists `y ∈ t` such that its preimage in `s` has no more than `n` elements. See also `Finset.exists_card_fiber_lt_of_card_lt_mul` for a stronger statement. -/ -theorem exists_card_fiber_le_of_card_le_mul (ht : t.Nonempty) (hn : s.card ≤ t.card * n) : - ∃ y ∈ t, (s.filter fun x => f x = y).card ≤ n := +theorem exists_card_fiber_le_of_card_le_mul (ht : t.Nonempty) (hn : #s ≤ #t * n) : + ∃ y ∈ t, #{x ∈ s | f x = y} ≤ n := exists_card_fiber_le_of_card_le_nsmul ht hn end Finset @@ -309,7 +306,7 @@ holds, so we have four theorems instead of eight. -/ version: there is a pigeonhole with the total weight of pigeons in it greater than `b` provided that the total number of pigeonholes times `b` is less than the total weight of all pigeons. -/ theorem exists_lt_sum_fiber_of_nsmul_lt_sum (hb : card β • b < ∑ x, w x) : - ∃ y, b < ∑ x ∈ univ.filter fun x => f x = y, w x := + ∃ y, b < ∑ x with f x = y, w x := let ⟨y, _, hy⟩ := exists_lt_sum_fiber_of_maps_to_of_nsmul_lt_sum (fun _ _ => mem_univ _) hb ⟨y, hy⟩ @@ -318,7 +315,7 @@ version: there is a pigeonhole with the total weight of pigeons in it greater th provided that the total number of pigeonholes times `b` is less than or equal to the total weight of all pigeons. -/ theorem exists_le_sum_fiber_of_nsmul_le_sum [Nonempty β] (hb : card β • b ≤ ∑ x, w x) : - ∃ y, b ≤ ∑ x ∈ univ.filter fun x => f x = y, w x := + ∃ y, b ≤ ∑ x with f x = y, w x := let ⟨y, _, hy⟩ := exists_le_sum_fiber_of_maps_to_of_nsmul_le_sum (fun _ _ => mem_univ _) univ_nonempty hb ⟨y, hy⟩ @@ -327,7 +324,7 @@ theorem exists_le_sum_fiber_of_nsmul_le_sum [Nonempty β] (hb : card β • b version: there is a pigeonhole with the total weight of pigeons in it less than `b` provided that the total number of pigeonholes times `b` is greater than the total weight of all pigeons. -/ theorem exists_sum_fiber_lt_of_sum_lt_nsmul (hb : ∑ x, w x < card β • b) : - ∃ y, ∑ x ∈ univ.filter fun x => f x = y, w x < b := + ∃ y, ∑ x with f x = y, w x < b := exists_lt_sum_fiber_of_nsmul_lt_sum (M := Mᵒᵈ) _ hb /-- The pigeonhole principle for finitely many pigeons of different weights, non-strict inequality @@ -335,7 +332,7 @@ version: there is a pigeonhole with the total weight of pigeons in it less than provided that the total number of pigeonholes times `b` is greater than or equal to the total weight of all pigeons. -/ theorem exists_sum_fiber_le_of_sum_le_nsmul [Nonempty β] (hb : ∑ x, w x ≤ card β • b) : - ∃ y, ∑ x ∈ univ.filter fun x => f x = y, w x ≤ b := + ∃ y, ∑ x with f x = y, w x ≤ b := exists_le_sum_fiber_of_nsmul_le_sum (M := Mᵒᵈ) _ hb end @@ -346,7 +343,7 @@ variable [LinearOrderedCommSemiring M] with at least as many pigeons as the ceiling of the average number of pigeons across all pigeonholes. -/ theorem exists_lt_card_fiber_of_nsmul_lt_card (hb : card β • b < card α) : - ∃ y : β, b < (univ.filter fun x => f x = y).card := + ∃ y : β, b < #{x | f x = y} := let ⟨y, _, h⟩ := exists_lt_card_fiber_of_nsmul_lt_card_of_maps_to (fun _ _ => mem_univ _) hb ⟨y, h⟩ @@ -359,14 +356,14 @@ More formally, given a function `f` between finite types `α` and `β` and a num `card β * n < card α`, there exists an element `y : β` such that its preimage has more than `n` elements. -/ theorem exists_lt_card_fiber_of_mul_lt_card (hn : card β * n < card α) : - ∃ y : β, n < (univ.filter fun x => f x = y).card := + ∃ y : β, n < #{x | f x = y} := exists_lt_card_fiber_of_nsmul_lt_card _ hn /-- The strong pigeonhole principle for finitely many pigeons and pigeonholes. There is a pigeonhole with at most as many pigeons as the floor of the average number of pigeons across all pigeonholes. -/ theorem exists_card_fiber_lt_of_card_lt_nsmul (hb : ↑(card α) < card β • b) : - ∃ y : β, ↑(univ.filter fun x => f x = y).card < b := + ∃ y : β, #{x | f x = y} < b := let ⟨y, _, h⟩ := Finset.exists_card_fiber_lt_of_card_lt_nsmul (f := f) hb ⟨y, h⟩ @@ -379,7 +376,7 @@ More formally, given a function `f` between finite types `α` and `β` and a num `card α < card β * n`, there exists an element `y : β` such that its preimage has less than `n` elements. -/ theorem exists_card_fiber_lt_of_card_lt_mul (hn : card α < card β * n) : - ∃ y : β, (univ.filter fun x => f x = y).card < n := + ∃ y : β, #{x | f x = y} < n := exists_card_fiber_lt_of_card_lt_nsmul _ hn /-- The strong pigeonhole principle for finitely many pigeons and pigeonholes. Given a function `f` @@ -387,7 +384,7 @@ between finite types `α` and `β` and a number `b` such that `card β • b ≤ element `y : β` such that its preimage has at least `b` elements. See also `Fintype.exists_lt_card_fiber_of_nsmul_lt_card` for a stronger statement. -/ theorem exists_le_card_fiber_of_nsmul_le_card [Nonempty β] (hb : card β • b ≤ card α) : - ∃ y : β, b ≤ (univ.filter fun x => f x = y).card := + ∃ y : β, b ≤ #{x | f x = y} := let ⟨y, _, h⟩ := exists_le_card_fiber_of_nsmul_le_card_of_maps_to (fun _ _ => mem_univ _) univ_nonempty hb ⟨y, h⟩ @@ -397,7 +394,7 @@ between finite types `α` and `β` and a number `n` such that `card β * n ≤ c element `y : β` such that its preimage has at least `n` elements. See also `Fintype.exists_lt_card_fiber_of_mul_lt_card` for a stronger statement. -/ theorem exists_le_card_fiber_of_mul_le_card [Nonempty β] (hn : card β * n ≤ card α) : - ∃ y : β, n ≤ (univ.filter fun x => f x = y).card := + ∃ y : β, n ≤ #{x | f x = y} := exists_le_card_fiber_of_nsmul_le_card _ hn /-- The strong pigeonhole principle for finitely many pigeons and pigeonholes. Given a function `f` @@ -405,7 +402,7 @@ between finite types `α` and `β` and a number `b` such that `card α ≤ card element `y : β` such that its preimage has at most `b` elements. See also `Fintype.exists_card_fiber_lt_of_card_lt_nsmul` for a stronger statement. -/ theorem exists_card_fiber_le_of_card_le_nsmul [Nonempty β] (hb : ↑(card α) ≤ card β • b) : - ∃ y : β, ↑(univ.filter fun x => f x = y).card ≤ b := + ∃ y : β, #{x | f x = y} ≤ b := let ⟨y, _, h⟩ := Finset.exists_card_fiber_le_of_card_le_nsmul univ_nonempty hb ⟨y, h⟩ @@ -414,7 +411,7 @@ between finite types `α` and `β` and a number `n` such that `card α ≤ card element `y : β` such that its preimage has at most `n` elements. See also `Fintype.exists_card_fiber_lt_of_card_lt_mul` for a stronger statement. -/ theorem exists_card_fiber_le_of_card_le_mul [Nonempty β] (hn : card α ≤ card β * n) : - ∃ y : β, (univ.filter fun x => f x = y).card ≤ n := + ∃ y : β, #{x | f x = y} ≤ n := exists_card_fiber_le_of_card_le_nsmul _ hn end Fintype diff --git a/Mathlib/Combinatorics/Schnirelmann.lean b/Mathlib/Combinatorics/Schnirelmann.lean index 266093dff23ab..537681ba1dae9 100644 --- a/Mathlib/Combinatorics/Schnirelmann.lean +++ b/Mathlib/Combinatorics/Schnirelmann.lean @@ -47,7 +47,7 @@ open Finset /-- The Schnirelmann density is defined as the infimum of |A ∩ {1, ..., n}| / n as n ranges over the positive naturals. -/ noncomputable def schnirelmannDensity (A : Set ℕ) [DecidablePred (· ∈ A)] : ℝ := - ⨅ n : {n : ℕ // 0 < n}, ((Ioc (0 : ℕ) n).filter (· ∈ A)).card / n + ⨅ n : {n : ℕ // 0 < n}, #{a ∈ Ioc 0 n | a ∈ A} / n section @@ -57,7 +57,7 @@ lemma schnirelmannDensity_nonneg : 0 ≤ schnirelmannDensity A := Real.iInf_nonneg (fun _ => by positivity) lemma schnirelmannDensity_le_div {n : ℕ} (hn : n ≠ 0) : - schnirelmannDensity A ≤ ((Ioc 0 n).filter (· ∈ A)).card / n := + schnirelmannDensity A ≤ #{a ∈ Ioc 0 n | a ∈ A} / n := ciInf_le ⟨0, fun _ ⟨_, hx⟩ => hx ▸ by positivity⟩ (⟨n, hn.bot_lt⟩ : {n : ℕ // 0 < n}) /-- @@ -65,7 +65,7 @@ For any natural `n`, the Schnirelmann density multiplied by `n` is bounded by `| Note this property fails for the natural density. -/ lemma schnirelmannDensity_mul_le_card_filter {n : ℕ} : - schnirelmannDensity A * n ≤ ((Ioc 0 n).filter (· ∈ A)).card := by + schnirelmannDensity A * n ≤ #{a ∈ Ioc 0 n | a ∈ A} := by rcases eq_or_ne n 0 with rfl | hn · simp exact (le_div_iff₀ (by positivity)).1 (schnirelmannDensity_le_div hn) @@ -78,7 +78,7 @@ We provide `n` explicitly here to make this lemma more easily usable in `apply` This lemma is analogous to `ciInf_le_of_le`. -/ lemma schnirelmannDensity_le_of_le {x : ℝ} (n : ℕ) (hn : n ≠ 0) - (hx : ((Ioc 0 n).filter (· ∈ A)).card / n ≤ x) : schnirelmannDensity A ≤ x := + (hx : #{a ∈ Ioc 0 n | a ∈ A} / n ≤ x) : schnirelmannDensity A ≤ x := (schnirelmannDensity_le_div hn).trans hx lemma schnirelmannDensity_le_one : schnirelmannDensity A ≤ 1 := @@ -96,7 +96,7 @@ lemma schnirelmannDensity_le_of_not_mem {k : ℕ} (hk : k ∉ A) : rw [← one_div, one_sub_div (Nat.cast_pos.2 hk').ne'] gcongr rw [← Nat.cast_pred hk', Nat.cast_le] - suffices (Ioc 0 k).filter (· ∈ A) ⊆ Ioo 0 k from (card_le_card this).trans_eq (by simp) + suffices {a ∈ Ioc 0 k | a ∈ A} ⊆ Ioo 0 k from (card_le_card this).trans_eq (by simp) rw [← Ioo_insert_right hk', filter_insert, if_neg hk] exact filter_subset _ _ @@ -138,16 +138,16 @@ lemma schnirelmannDensity_eq_one_iff_of_zero_mem (hA : 0 ∈ A) : exact Set.subset_univ {0}ᶜ lemma le_schnirelmannDensity_iff {x : ℝ} : - x ≤ schnirelmannDensity A ↔ ∀ n : ℕ, 0 < n → x ≤ ((Ioc 0 n).filter (· ∈ A)).card / n := + x ≤ schnirelmannDensity A ↔ ∀ n : ℕ, 0 < n → x ≤ #{a ∈ Ioc 0 n | a ∈ A} / n := (le_ciInf_iff ⟨0, fun _ ⟨_, hx⟩ => hx ▸ by positivity⟩).trans Subtype.forall lemma schnirelmannDensity_lt_iff {x : ℝ} : - schnirelmannDensity A < x ↔ ∃ n : ℕ, 0 < n ∧ ((Ioc 0 n).filter (· ∈ A)).card / n < x := by + schnirelmannDensity A < x ↔ ∃ n : ℕ, 0 < n ∧ #{a ∈ Ioc 0 n | a ∈ A} / n < x := by rw [← not_le, le_schnirelmannDensity_iff]; simp lemma schnirelmannDensity_le_iff_forall {x : ℝ} : schnirelmannDensity A ≤ x ↔ - ∀ ε : ℝ, 0 < ε → ∃ n : ℕ, 0 < n ∧ ((Ioc 0 n).filter (· ∈ A)).card / n < x + ε := by + ∀ ε : ℝ, 0 < ε → ∃ n : ℕ, 0 < n ∧ #{a ∈ Ioc 0 n | a ∈ A} / n < x + ε := by rw [le_iff_forall_pos_lt_add] simp only [schnirelmannDensity_lt_iff] @@ -175,7 +175,7 @@ If the Schnirelmann density is `0`, there is a positive natural for which Note this cannot be improved to `∃ᶠ n : ℕ in atTop`, as can be seen by `A = {1}ᶜ`. -/ lemma exists_of_schnirelmannDensity_eq_zero {ε : ℝ} (hε : 0 < ε) (hA : schnirelmannDensity A = 0) : - ∃ n, 0 < n ∧ ((Ioc 0 n).filter (· ∈ A)).card / n < ε := by + ∃ n, 0 < n ∧ #{a ∈ Ioc 0 n | a ∈ A} / n < ε := by by_contra! h rw [← le_schnirelmannDensity_iff] at h linarith @@ -193,7 +193,7 @@ lemma schnirelmannDensity_finset (A : Finset ℕ) : schnirelmannDensity A = 0 := wlog hε₁ : ε ≤ 1 generalizing ε · obtain ⟨n, hn, hn'⟩ := this 1 zero_lt_one le_rfl exact ⟨n, hn, hn'.trans_le (le_of_not_le hε₁)⟩ - let n : ℕ := ⌊A.card / ε⌋₊ + 1 + let n : ℕ := ⌊#A / ε⌋₊ + 1 have hn : 0 < n := Nat.succ_pos _ use n, hn rw [div_lt_iff₀ (Nat.cast_pos.2 hn), ← div_lt_iff₀' hε, Nat.cast_add_one] @@ -236,7 +236,7 @@ lemma schnirelmannDensity_setOf_mod_eq_one {m : ℕ} (hm : m ≠ 1) : rw [le_schnirelmannDensity_iff] intro n hn simp only [Set.mem_setOf_eq] - have : (Icc 0 ((n - 1) / m)).image (· * m + 1) ⊆ (Ioc 0 n).filter (· % m = 1) := by + have : (Icc 0 ((n - 1) / m)).image (· * m + 1) ⊆ {x ∈ Ioc 0 n | x % m = 1} := by simp only [subset_iff, mem_image, forall_exists_index, mem_filter, mem_Ioc, mem_Icc, and_imp] rintro _ y _ hy' rfl have hm : 2 ≤ m := hm.lt_of_le' hm' diff --git a/Mathlib/Combinatorics/SetFamily/AhlswedeZhang.lean b/Mathlib/Combinatorics/SetFamily/AhlswedeZhang.lean index 73740d7fb77d3..a7d1c6345f5ca 100644 --- a/Mathlib/Combinatorics/SetFamily/AhlswedeZhang.lean +++ b/Mathlib/Combinatorics/SetFamily/AhlswedeZhang.lean @@ -69,11 +69,11 @@ private lemma binomial_sum_eq (h : n < m) : ring private lemma Fintype.sum_div_mul_card_choose_card : - ∑ s : Finset α, (card α / ((card α - s.card) * (card α).choose s.card) : ℚ) = + ∑ s : Finset α, (card α / ((card α - #s) * (card α).choose #s) : ℚ) = card α * ∑ k ∈ range (card α), (↑k)⁻¹ + 1 := by rw [← powerset_univ, powerset_card_disjiUnion, sum_disjiUnion] have : ∀ {x : ℕ}, ∀ s ∈ powersetCard x (univ : Finset α), - (card α / ((card α - Finset.card s) * (card α).choose (Finset.card s)) : ℚ) = + (card α / ((card α - #s) * (card α).choose #s) : ℚ) = card α / ((card α - x) * (card α).choose x) := by intros n s hs rw [mem_powersetCard_univ.1 hs] @@ -103,8 +103,7 @@ section SemilatticeSup variable [SemilatticeSup α] [SemilatticeSup β] [BoundedOrder β] {s t : Finset α} {a b : α} -private lemma sup_aux [@DecidableRel α (· ≤ ·)] : - a ∈ lowerClosure s → (s.filter fun b ↦ a ≤ b).Nonempty := +private lemma sup_aux [@DecidableRel α (· ≤ ·)] : a ∈ lowerClosure s → {b ∈ s | a ≤ b}.Nonempty := fun ⟨b, hb, hab⟩ ↦ ⟨b, mem_filter.2 ⟨hb, hab⟩⟩ private lemma lower_aux [DecidableEq α] : @@ -115,10 +114,10 @@ variable [@DecidableRel α (· ≤ ·)] [OrderTop α] /-- The supremum of the elements of `s` less than `a` if there are some, otherwise `⊤`. -/ def truncatedSup (s : Finset α) (a : α) : α := - if h : a ∈ lowerClosure s then (s.filter fun b ↦ a ≤ b).sup' (sup_aux h) id else ⊤ + if h : a ∈ lowerClosure s then {b ∈ s | a ≤ b}.sup' (sup_aux h) id else ⊤ lemma truncatedSup_of_mem (h : a ∈ lowerClosure s) : - truncatedSup s a = (s.filter fun b ↦ a ≤ b).sup' (sup_aux h) id := dif_pos h + truncatedSup s a = {b ∈ s | a ≤ b}.sup' (sup_aux h) id := dif_pos h lemma truncatedSup_of_not_mem (h : a ∉ lowerClosure s) : truncatedSup s a = ⊤ := dif_neg h @@ -177,7 +176,7 @@ section SemilatticeInf variable [SemilatticeInf α] [SemilatticeInf β] [BoundedOrder β] [@DecidableRel β (· ≤ ·)] {s t : Finset α} {a : α} -private lemma inf_aux [@DecidableRel α (· ≤ ·)]: a ∈ upperClosure s → (s.filter (· ≤ a)).Nonempty := +private lemma inf_aux [@DecidableRel α (· ≤ ·)]: a ∈ upperClosure s → {b ∈ s | b ≤ a}.Nonempty := fun ⟨b, hb, hab⟩ ↦ ⟨b, mem_filter.2 ⟨hb, hab⟩⟩ private lemma upper_aux [DecidableEq α] : @@ -188,10 +187,10 @@ variable [@DecidableRel α (· ≤ ·)] [BoundedOrder α] /-- The infimum of the elements of `s` less than `a` if there are some, otherwise `⊥`. -/ def truncatedInf (s : Finset α) (a : α) : α := - if h : a ∈ upperClosure s then (s.filter (· ≤ a)).inf' (inf_aux h) id else ⊥ + if h : a ∈ upperClosure s then {b ∈ s | b ≤ a}.inf' (inf_aux h) id else ⊥ lemma truncatedInf_of_mem (h : a ∈ upperClosure s) : - truncatedInf s a = (s.filter (· ≤ a)).inf' (inf_aux h) id := dif_pos h + truncatedInf s a = {b ∈ s | b ≤ a}.inf' (inf_aux h) id := dif_pos h lemma truncatedInf_of_not_mem (h : a ∉ upperClosure s) : truncatedInf s a = ⊥ := dif_neg h @@ -297,8 +296,8 @@ end BooleanAlgebra variable [DecidableEq α] [Fintype α] lemma card_truncatedSup_union_add_card_truncatedSup_infs (𝒜 ℬ : Finset (Finset α)) (s : Finset α) : - (truncatedSup (𝒜 ∪ ℬ) s).card + (truncatedSup (𝒜 ⊼ ℬ) s).card = - (truncatedSup 𝒜 s).card + (truncatedSup ℬ s).card := by + #(truncatedSup (𝒜 ∪ ℬ) s) + #(truncatedSup (𝒜 ⊼ ℬ) s) = + #(truncatedSup 𝒜 s) + #(truncatedSup ℬ s) := by by_cases h𝒜 : s ∈ lowerClosure (𝒜 : Set <| Finset α) <;> by_cases hℬ : s ∈ lowerClosure (ℬ : Set <| Finset α) · rw [truncatedSup_union h𝒜 hℬ, truncatedSup_infs h𝒜 hℬ] @@ -311,8 +310,8 @@ lemma card_truncatedSup_union_add_card_truncatedSup_infs (𝒜 ℬ : Finset (Fin truncatedSup_union_of_not_mem h𝒜 hℬ, truncatedSup_infs_of_not_mem fun h ↦ h𝒜 h.1] lemma card_truncatedInf_union_add_card_truncatedInf_sups (𝒜 ℬ : Finset (Finset α)) (s : Finset α) : - (truncatedInf (𝒜 ∪ ℬ) s).card + (truncatedInf (𝒜 ⊻ ℬ) s).card = - (truncatedInf 𝒜 s).card + (truncatedInf ℬ s).card := by + #(truncatedInf (𝒜 ∪ ℬ) s) + #(truncatedInf (𝒜 ⊻ ℬ) s) = + #(truncatedInf 𝒜 s) + #(truncatedInf ℬ s) := by by_cases h𝒜 : s ∈ upperClosure (𝒜 : Set <| Finset α) <;> by_cases hℬ : s ∈ upperClosure (ℬ : Set <| Finset α) · rw [truncatedInf_union h𝒜 hℬ, truncatedInf_sups h𝒜 hℬ] @@ -335,12 +334,12 @@ variable {α : Type*} [Fintype α] [DecidableEq α] {𝒜 ℬ : Finset (Finset /-- Weighted sum of the size of the truncated infima of a set family. Relevant to the Ahlswede-Zhang identity. -/ def infSum (𝒜 : Finset (Finset α)) : ℚ := - ∑ s, (truncatedInf 𝒜 s).card / (s.card * (card α).choose s.card) + ∑ s, #(truncatedInf 𝒜 s) / (#s * (card α).choose #s) /-- Weighted sum of the size of the truncated suprema of a set family. Relevant to the Ahlswede-Zhang identity. -/ def supSum (𝒜 : Finset (Finset α)) : ℚ := - ∑ s, (truncatedSup 𝒜 s).card / ((card α - s.card) * (card α).choose s.card) + ∑ s, #(truncatedSup 𝒜 s) / ((card α - #s) * (card α).choose #s) lemma supSum_union_add_supSum_infs (𝒜 ℬ : Finset (Finset α)) : supSum (𝒜 ∪ ℬ) + supSum (𝒜 ⊼ ℬ) = supSum 𝒜 + supSum ℬ := by @@ -357,9 +356,9 @@ lemma infSum_union_add_infSum_sups (𝒜 ℬ : Finset (Finset α)) : simp lemma IsAntichain.le_infSum (h𝒜 : IsAntichain (· ⊆ ·) (𝒜 : Set (Finset α))) (h𝒜₀ : ∅ ∉ 𝒜) : - ∑ s ∈ 𝒜, ((card α).choose s.card : ℚ)⁻¹ ≤ infSum 𝒜 := by + ∑ s ∈ 𝒜, ((card α).choose #s : ℚ)⁻¹ ≤ infSum 𝒜 := by calc - _ = ∑ s ∈ 𝒜, (truncatedInf 𝒜 s).card / (s.card * (card α).choose s.card : ℚ) := ?_ + _ = ∑ s ∈ 𝒜, #(truncatedInf 𝒜 s) / (#s * (card α).choose #s : ℚ) := ?_ _ ≤ _ := sum_le_univ_sum_of_nonneg fun s ↦ by positivity refine sum_congr rfl fun s hs ↦ ?_ rw [truncatedInf_of_isAntichain h𝒜 hs, div_mul_cancel_left₀] @@ -371,8 +370,8 @@ variable [Nonempty α] @[simp] lemma supSum_singleton (hs : s ≠ univ) : supSum ({s} : Finset (Finset α)) = card α * ∑ k ∈ range (card α), (k : ℚ)⁻¹ := by have : ∀ t : Finset α, - (card α - (truncatedSup {s} t).card : ℚ) / ((card α - t.card) * (card α).choose t.card) = - if t ⊆ s then (card α - s.card : ℚ) / ((card α - t.card) * (card α).choose t.card) else 0 := by + (card α - #(truncatedSup {s} t) : ℚ) / ((card α - #t) * (card α).choose #t) = + if t ⊆ s then (card α - #s : ℚ) / ((card α - #t) * (card α).choose #t) else 0 := by rintro t simp_rw [truncatedSup_singleton, le_iff_subset] split_ifs <;> simp [card_univ] @@ -382,7 +381,7 @@ variable [Nonempty α] sum_powerset, ← binomial_sum_eq ((card_lt_iff_ne_univ _).2 hs), eq_comm] refine sum_congr rfl fun n _ ↦ ?_ rw [mul_div_assoc, ← nsmul_eq_mul] - exact sum_powersetCard n s fun m ↦ (card α - s.card : ℚ) / ((card α - m) * (card α).choose m) + exact sum_powersetCard n s fun m ↦ (card α - #s : ℚ) / ((card α - m) * (card α).choose m) /-- The **Ahlswede-Zhang Identity**. -/ lemma infSum_compls_add_supSum (𝒜 : Finset (Finset α)) : diff --git a/Mathlib/Combinatorics/SetFamily/CauchyDavenport.lean b/Mathlib/Combinatorics/SetFamily/CauchyDavenport.lean index aed46e2c9d07d..a08ee82af2783 100644 --- a/Mathlib/Combinatorics/SetFamily/CauchyDavenport.lean +++ b/Mathlib/Combinatorics/SetFamily/CauchyDavenport.lean @@ -71,27 +71,24 @@ variable [Group α] [DecidableEq α] {x y : Finset α × Finset α} {s t : Finse * or `|s₁ + t₁| = |s₂ + t₂|` and `|s₂| + |t₂| < |s₁| + |t₁|` * or `|s₁ + t₁| = |s₂ + t₂|` and `|s₁| + |t₁| = |s₂| + |t₂|` and `|s₁| < |s₂|`."] private def DevosMulRel : Finset α × Finset α → Finset α × Finset α → Prop := - Prod.Lex (· < ·) (Prod.Lex (· > ·) (· < ·)) on fun x ↦ - ((x.1 * x.2).card, x.1.card + x.2.card, x.1.card) + Prod.Lex (· < ·) (Prod.Lex (· > ·) (· < ·)) on fun x ↦ (#(x.1 * x.2), #x.1 + #x.2, #x.1) @[to_additive] private lemma devosMulRel_iff : DevosMulRel x y ↔ - (x.1 * x.2).card < (y.1 * y.2).card ∨ - (x.1 * x.2).card = (y.1 * y.2).card ∧ y.1.card + y.2.card < x.1.card + x.2.card ∨ - (x.1 * x.2).card = (y.1 * y.2).card ∧ - x.1.card + x.2.card = y.1.card + y.2.card ∧ x.1.card < y.1.card := by + #(x.1 * x.2) < #(y.1 * y.2) ∨ + #(x.1 * x.2) = #(y.1 * y.2) ∧ #y.1 + #y.2 < #x.1 + #x.2 ∨ + #(x.1 * x.2) = #(y.1 * y.2) ∧ #x.1 + #x.2 = #y.1 + #y.2 ∧ #x.1 < #y.1 := by simp [DevosMulRel, Prod.lex_iff, and_or_left] @[to_additive] -private lemma devosMulRel_of_le (mul : (x.1 * x.2).card ≤ (y.1 * y.2).card) - (hadd : y.1.card + y.2.card < x.1.card + x.2.card) : DevosMulRel x y := +private lemma devosMulRel_of_le (mul : #(x.1 * x.2) ≤ #(y.1 * y.2)) + (hadd : #y.1 + #y.2 < #x.1 + #x.2) : DevosMulRel x y := devosMulRel_iff.2 <| mul.lt_or_eq.imp_right fun h ↦ Or.inl ⟨h, hadd⟩ @[to_additive] -private lemma devosMulRel_of_le_of_le (mul : (x.1 * x.2).card ≤ (y.1 * y.2).card) - (hadd : y.1.card + y.2.card ≤ x.1.card + x.2.card) (hone : x.1.card < y.1.card) : - DevosMulRel x y := +private lemma devosMulRel_of_le_of_le (mul : #(x.1 * x.2) ≤ #(y.1 * y.2)) + (hadd : #y.1 + #y.2 ≤ #x.1 + #x.2) (hone : #x.1 < #y.1) : DevosMulRel x y := devosMulRel_iff.2 <| mul.lt_or_eq.imp_right fun h ↦ hadd.gt_or_eq.imp (And.intro h) fun h' ↦ ⟨h, h', hone⟩ @@ -113,20 +110,20 @@ subgroup. -/ `s + t` is lower-bounded by `|s| + |t| - 1` unless this quantity is greater than the size of the smallest subgroup."] lemma Finset.min_le_card_mul (hs : s.Nonempty) (ht : t.Nonempty) : - min (minOrder α) ↑(s.card + t.card - 1) ≤ (s * t).card := by + min (minOrder α) ↑(#s + #t - 1) ≤ #(s * t) := by -- Set up the induction on `x := (s, t)` along the `DevosMulRel` relation. set x := (s, t) with hx clear_value x simp only [Prod.ext_iff] at hx obtain ⟨rfl, rfl⟩ := hx refine wellFoundedOn_devosMulRel.induction (P := fun x : Finset α × Finset α ↦ - min (minOrder α) ↑(card x.1 + card x.2 - 1) ≤ card (x.1 * x.2)) ⟨hs, ht⟩ ?_ + min (minOrder α) ↑(#x.1 + #x.2 - 1) ≤ #(x.1 * x.2)) ⟨hs, ht⟩ ?_ clear! x rintro ⟨s, t⟩ ⟨hs, ht⟩ ih simp only [min_le_iff, tsub_le_iff_right, Prod.forall, Set.mem_setOf_eq, and_imp, Nat.cast_le] at * - -- If `t.card < s.card`, we're done by the induction hypothesis on `(t⁻¹, s⁻¹)`. - obtain hts | hst := lt_or_le t.card s.card + -- If `#t < #s`, we're done by the induction hypothesis on `(t⁻¹, s⁻¹)`. + obtain hts | hst := lt_or_le #t #s · simpa only [← mul_inv_rev, add_comm, card_inv] using ih _ _ ht.inv hs.inv (devosMulRel_iff.2 <| Or.inr <| Or.inr <| by @@ -159,7 +156,7 @@ lemma Finset.min_le_card_mul (hs : s.Nonempty) (ht : t.Nonempty) : simp [-coe_smul_finset] -- Else, we can transform `s`, `t` to `s'`, `t'` and `s''`, `t''`, such that one of `(s', t')` and -- `(s'', t'')` is strictly smaller than `(s, t)` according to `DevosMulRel`. - replace hsg : (s ∩ op g • s).card < s.card := card_lt_card ⟨inter_subset_left, fun h ↦ + replace hsg : #(s ∩ op g • s) < #s := card_lt_card ⟨inter_subset_left, fun h ↦ hsg <| eq_of_superset_of_card_ge (h.trans inter_subset_right) (card_smul_finset _ _).le⟩ replace aux1 := card_mono <| mulETransformLeft.fst_mul_snd_subset g (s, t) replace aux2 := card_mono <| mulETransformRight.fst_mul_snd_subset g (s, t) @@ -184,7 +181,7 @@ by `|s| + |t| - 1`. -/ @[to_additive "The **Cauchy-Davenport theorem** for torsion-free groups. The size of `s + t` is lower-bounded by `|s| + |t| - 1`."] lemma Monoid.IsTorsionFree.card_add_card_sub_one_le_card_mul (h : IsTorsionFree α) - (hs : s.Nonempty) (ht : t.Nonempty) : s.card + t.card - 1 ≤ (s * t).card := by + (hs : s.Nonempty) (ht : t.Nonempty) : #s + #t - 1 ≤ #(s * t) := by simpa only [h.minOrder, min_eq_right, le_top, Nat.cast_le] using Finset.min_le_card_mul hs ht end General @@ -194,7 +191,7 @@ end General /-- The **Cauchy-Davenport Theorem**. If `s`, `t` are nonempty sets in $$ℤ/pℤ$$, then the size of `s + t` is lower-bounded by `|s| + |t| - 1`, unless this quantity is greater than `p`. -/ lemma ZMod.min_le_card_add {p : ℕ} (hp : p.Prime) {s t : Finset (ZMod p)} (hs : s.Nonempty) - (ht : t.Nonempty) : min p (s.card + t.card - 1) ≤ (s + t).card := by + (ht : t.Nonempty) : min p (#s + #t - 1) ≤ #(s + t) := by simpa only [ZMod.minOrder_of_prime hp, min_le_iff, Nat.cast_le] using Finset.min_le_card_add hs ht /-! ### Linearly ordered cancellative semigroups -/ @@ -206,7 +203,7 @@ lemma ZMod.min_le_card_add {p : ℕ} (hp : p.Prime) {s t : Finset (ZMod p)} (hs `s + t` is lower-bounded by `|s| + |t| - 1`."] lemma Finset.card_add_card_sub_one_le_card_mul [LinearOrder α] [Semigroup α] [IsCancelMul α] [CovariantClass α α (· * ·) (· ≤ ·)] [CovariantClass α α (swap (· * ·)) (· ≤ ·)] - {s t : Finset α} (hs : s.Nonempty) (ht : t.Nonempty) : s.card + t.card - 1 ≤ (s * t).card := by + {s t : Finset α} (hs : s.Nonempty) (ht : t.Nonempty) : #s + #t - 1 ≤ #(s * t) := by suffices s * {t.min' ht} ∩ ({s.max' hs} * t) = {s.max' hs * t.min' ht} by rw [← card_singleton_mul t (s.max' hs), ← card_mul_singleton s (t.min' ht), ← card_union_add_card_inter, ← card_singleton _, ← this, Nat.add_sub_cancel] diff --git a/Mathlib/Combinatorics/SetFamily/Compression/Down.lean b/Mathlib/Combinatorics/SetFamily/Compression/Down.lean index a93982a2764cf..6a0dbfd4822da 100644 --- a/Mathlib/Combinatorics/SetFamily/Compression/Down.lean +++ b/Mathlib/Combinatorics/SetFamily/Compression/Down.lean @@ -40,13 +40,12 @@ variable {α : Type*} [DecidableEq α] {𝒜 : Finset (Finset α)} {s : Finset namespace Finset /-- Elements of `𝒜` that do not contain `a`. -/ -def nonMemberSubfamily (a : α) (𝒜 : Finset (Finset α)) : Finset (Finset α) := - 𝒜.filter fun s => a ∉ s +def nonMemberSubfamily (a : α) (𝒜 : Finset (Finset α)) : Finset (Finset α) := {s ∈ 𝒜 | a ∉ s} /-- Image of the elements of `𝒜` which contain `a` under removing `a`. Finsets that do not contain `a` such that `insert a s ∈ 𝒜`. -/ def memberSubfamily (a : α) (𝒜 : Finset (Finset α)) : Finset (Finset α) := - (𝒜.filter fun s => a ∈ s).image fun s => erase s a + {s ∈ 𝒜 | a ∈ s}.image fun s => erase s a @[simp] theorem mem_nonMemberSubfamily : s ∈ 𝒜.nonMemberSubfamily a ↔ s ∈ 𝒜 ∧ a ∉ s := by @@ -79,7 +78,7 @@ theorem memberSubfamily_union (a : α) (𝒜 ℬ : Finset (Finset α)) : simp_rw [memberSubfamily, filter_union, image_union] theorem card_memberSubfamily_add_card_nonMemberSubfamily (a : α) (𝒜 : Finset (Finset α)) : - (𝒜.memberSubfamily a).card + (𝒜.nonMemberSubfamily a).card = 𝒜.card := by + #(𝒜.memberSubfamily a) + #(𝒜.nonMemberSubfamily a) = #𝒜 := by rw [memberSubfamily, nonMemberSubfamily, card_image_of_injOn] · conv_rhs => rw [← filter_card_add_filter_neg_card_eq_card (fun s => (a ∈ s))] · apply (erase_injOn' _).mono @@ -136,7 +135,7 @@ lemma memberSubfamily_image_insert (h𝒜 : ∀ s ∈ 𝒜, a ∉ s) : (ne_of_mem_of_not_mem' (mem_insert_self _ _) (not_mem_erase _ _)).symm] lemma image_insert_memberSubfamily (𝒜 : Finset (Finset α)) (a : α) : - (𝒜.memberSubfamily a).image (insert a) = 𝒜.filter (a ∈ ·) := by + (𝒜.memberSubfamily a).image (insert a) = {s ∈ 𝒜 | a ∈ s} := by ext s simp only [mem_memberSubfamily, mem_image, mem_filter] refine ⟨?_, fun ⟨hs, ha⟩ ↦ ⟨erase s a, ⟨?_, not_mem_erase _ _⟩, insert_erase ha⟩⟩ @@ -182,7 +181,7 @@ it suffices to prove it for * `ℬ ∪ 𝒞` assuming the property for `ℬ` and `𝒞`, where `a` is an element of the ground type and `ℬ`is a family of finsets not containing `a` and `𝒞` a family of finsets containing `a`. Note that instead of giving `ℬ` and `𝒞`, the `subfamily` case gives you `𝒜 = ℬ ∪ 𝒞`, so that - `ℬ = 𝒜.filter (a ∉ ·)` and `𝒞 = 𝒜.filter (a ∈ ·)`. + `ℬ = {s ∈ 𝒜 | a ∉ s}` and `𝒞 = {s ∈ 𝒜 | a ∈ s}`. This is a way of formalising induction on `n` where `𝒜` is a finset family on `n` elements. @@ -193,7 +192,7 @@ protected lemma family_induction_on {p : Finset (Finset α) → Prop} (image_insert : ∀ (a : α) ⦃𝒜 : Finset (Finset α)⦄, (∀ s ∈ 𝒜, a ∉ s) → p 𝒜 → p (𝒜.image <| insert a)) (subfamily : ∀ (a : α) ⦃𝒜 : Finset (Finset α)⦄, - p (𝒜.filter (a ∉ ·)) → p (𝒜.filter (a ∈ ·)) → p 𝒜) : p 𝒜 := by + p {s ∈ 𝒜 | a ∉ s} → p {s ∈ 𝒜 | a ∈ s} → p 𝒜) : p 𝒜 := by refine memberFamily_induction_on 𝒜 empty singleton_empty fun a 𝒜 h𝒜₀ h𝒜₁ ↦ subfamily a h𝒜₀ ?_ rw [← image_insert_memberSubfamily] exact image_insert _ (by simp) h𝒜₁ @@ -208,11 +207,8 @@ namespace Down /-- `a`-down-compressing `𝒜` means removing `a` from the elements of `𝒜` that contain it, when the resulting Finset is not already in `𝒜`. -/ def compression (a : α) (𝒜 : Finset (Finset α)) : Finset (Finset α) := - (𝒜.filter fun s => erase s a ∈ 𝒜).disjUnion - ((𝒜.image fun s => erase s a).filter fun s => s ∉ 𝒜) <| - disjoint_left.2 fun s h₁ h₂ => by - have := (mem_filter.1 h₂).2 - exact this (mem_filter.1 h₁).1 + {s ∈ 𝒜 | erase s a ∈ 𝒜}.disjUnion {s ∈ 𝒜.image fun s ↦ erase s a | s ∉ 𝒜} <| + disjoint_left.2 fun _s h₁ h₂ ↦ (mem_filter.1 h₂).2 (mem_filter.1 h₁).1 @[inherit_doc] scoped[FinsetFamily] notation "𝓓 " => Down.compression @@ -258,7 +254,7 @@ theorem compression_idem (a : α) (𝒜 : Finset (Finset α)) : 𝓓 a (𝓓 a /-- Down-compressing a family doesn't change its size. -/ @[simp] -theorem card_compression (a : α) (𝒜 : Finset (Finset α)) : (𝓓 a 𝒜).card = 𝒜.card := by +theorem card_compression (a : α) (𝒜 : Finset (Finset α)) : #(𝓓 a 𝒜) = #𝒜 := by rw [compression, card_disjUnion, filter_image, card_image_of_injOn ((erase_injOn' _).mono fun s hs => _), ← card_union_of_disjoint] · conv_rhs => rw [← filter_union_filter_neg_eq (fun s => (erase s a ∈ 𝒜)) 𝒜] diff --git a/Mathlib/Combinatorics/SetFamily/Compression/UV.lean b/Mathlib/Combinatorics/SetFamily/Compression/UV.lean index 8bfa4e268eb8a..537a0a3db2b6f 100644 --- a/Mathlib/Combinatorics/SetFamily/Compression/UV.lean +++ b/Mathlib/Combinatorics/SetFamily/Compression/UV.lean @@ -116,7 +116,7 @@ variable [DecidableEq α] /-- To UV-compress a set family, we compress each of its elements, except that we don't want to reduce the cardinality, so we keep all elements whose compression is already present. -/ def compression (u v : α) (s : Finset α) := - (s.filter (compress u v · ∈ s)) ∪ (s.image <| compress u v).filter (· ∉ s) + {a ∈ s | compress u v a ∈ s} ∪ {a ∈ s.image <| compress u v | a ∉ s} @[inherit_doc] scoped[FinsetFamily] notation "𝓒 " => UV.compression @@ -128,7 +128,7 @@ def IsCompressed (u v : α) (s : Finset α) := 𝓒 u v s = s /-- UV-compression is injective on the sets that are not UV-compressed. -/ -theorem compress_injOn : Set.InjOn (compress u v) ↑(s.filter (compress u v · ∉ s)) := by +theorem compress_injOn : Set.InjOn (compress u v) ↑{a ∈ s | compress u v a ∉ s} := by intro a ha b hb hab rw [mem_coe, mem_filter] at ha hb rw [compress] at ha hab @@ -162,7 +162,7 @@ theorem compression_self (u : α) (s : Finset α) : 𝓒 u u s = s := by theorem isCompressed_self (u : α) (s : Finset α) : IsCompressed u u s := compression_self u s theorem compress_disjoint : - Disjoint (s.filter (compress u v · ∈ s)) ((s.image <| compress u v).filter (· ∉ s)) := + Disjoint {a ∈ s | compress u v a ∈ s} {a ∈ s.image <| compress u v | a ∉ s} := disjoint_left.2 fun _a ha₁ ha₂ ↦ (mem_filter.1 ha₂).2 (mem_filter.1 ha₁).1 theorem compress_mem_compression (ha : a ∈ s) : compress u v a ∈ 𝓒 u v s := by @@ -184,14 +184,14 @@ theorem compress_mem_compression_of_mem_compression (ha : a ∈ 𝓒 u v s) : /-- Compressing a family is idempotent. -/ @[simp] theorem compression_idem (u v : α) (s : Finset α) : 𝓒 u v (𝓒 u v s) = 𝓒 u v s := by - have h : filter (compress u v · ∉ 𝓒 u v s) (𝓒 u v s) = ∅ := + have h : {a ∈ 𝓒 u v s | compress u v a ∉ 𝓒 u v s} = ∅ := filter_false_of_mem fun a ha h ↦ h <| compress_mem_compression_of_mem_compression ha rw [compression, filter_image, h, image_empty, ← h] exact filter_union_filter_neg_eq _ (compression u v s) /-- Compressing a family doesn't change its size. -/ @[simp] -theorem card_compression (u v : α) (s : Finset α) : (𝓒 u v s).card = s.card := by +theorem card_compression (u v : α) (s : Finset α) : #(𝓒 u v s) = #s := by rw [compression, card_union_of_disjoint compress_disjoint, filter_image, card_image_of_injOn compress_injOn, ← card_union_of_disjoint (disjoint_filter_filter_neg s _ _), filter_union_filter_neg_eq] @@ -268,14 +268,14 @@ open FinsetFamily variable [DecidableEq α] {𝒜 : Finset (Finset α)} {u v a : Finset α} {r : ℕ} /-- Compressing a finset doesn't change its size. -/ -theorem card_compress (huv : u.card = v.card) (a : Finset α) : (compress u v a).card = a.card := by +theorem card_compress (huv : #u = #v) (a : Finset α) : #(compress u v a) = #a := by unfold compress split_ifs with h · rw [card_sdiff (h.2.trans le_sup_left), sup_eq_union, card_union_of_disjoint h.1.symm, huv, add_tsub_cancel_right] · rfl -lemma _root_.Set.Sized.uvCompression (huv : u.card = v.card) (h𝒜 : (𝒜 : Set (Finset α)).Sized r) : +lemma _root_.Set.Sized.uvCompression (huv : #u = #v) (h𝒜 : (𝒜 : Set (Finset α)).Sized r) : (𝓒 u v 𝒜 : Set (Finset α)).Sized r := by simp_rw [Set.Sized, mem_coe, mem_compression] rintro s (hs | ⟨huvt, t, ht, rfl⟩) @@ -397,7 +397,7 @@ such that `𝒜` is `(u.erase x, v.erase y)`-compressed. This is the key UV-comp Kruskal-Katona. -/ theorem card_shadow_compression_le (u v : Finset α) (huv : ∀ x ∈ u, ∃ y ∈ v, IsCompressed (u.erase x) (v.erase y) 𝒜) : - (∂ (𝓒 u v 𝒜)).card ≤ (∂ 𝒜).card := + #(∂ (𝓒 u v 𝒜)) ≤ #(∂ 𝒜) := (card_le_card <| shadow_compression_subset_compression_shadow _ _ huv).trans (card_compression _ _ _).le diff --git a/Mathlib/Combinatorics/SetFamily/FourFunctions.lean b/Mathlib/Combinatorics/SetFamily/FourFunctions.lean index df03a09a30705..6f6dbafc9c630 100644 --- a/Mathlib/Combinatorics/SetFamily/FourFunctions.lean +++ b/Mathlib/Combinatorics/SetFamily/FourFunctions.lean @@ -87,7 +87,7 @@ private lemma ineq [ExistsAddOfLE β] {a₀ a₁ b₀ b₁ c₀ c₁ d₀ d₁ : _ = (c₀ * d₁ + c₁ * d₀) * (c₀ * d₁) := by ring private def collapse (𝒜 : Finset (Finset α)) (a : α) (f : Finset α → β) (s : Finset α) : β := - ∑ t ∈ 𝒜.filter fun t ↦ t.erase a = s, f t + ∑ t ∈ 𝒜 with t.erase a = s, f t private lemma erase_eq_iff (hs : a ∉ s) : t.erase a = s ↔ t = s ∨ t = insert a s := by by_cases ht : a ∈ t <;> @@ -95,7 +95,7 @@ private lemma erase_eq_iff (hs : a ∉ s) : t.erase a = s ↔ t = s ∨ t = inse aesop private lemma filter_collapse_eq (ha : a ∉ s) (𝒜 : Finset (Finset α)) : - (𝒜.filter fun t ↦ t.erase a = s) = + {t ∈ 𝒜 | t.erase a = s} = if s ∈ 𝒜 then (if insert a s ∈ 𝒜 then {s, insert a s} else {s}) else @@ -304,7 +304,7 @@ lemma four_functions_theorem [DecidableEq α] (h₁ : 0 ≤ f₁) (h₂ : 0 ≤ /-- An inequality of Daykin. Interestingly, any lattice in which this inequality holds is distributive. -/ lemma Finset.le_card_infs_mul_card_sups [DecidableEq α] (s t : Finset α) : - s.card * t.card ≤ (s ⊼ t).card * (s ⊻ t).card := by + #s * #t ≤ #(s ⊼ t) * #(s ⊻ t) := by simpa using four_functions_theorem (1 : α → ℕ) 1 1 1 zero_le_one zero_le_one zero_le_one zero_le_one (fun _ _ ↦ le_rfl) s t @@ -353,7 +353,7 @@ variable [DecidableEq α] [GeneralizedBooleanAlgebra α] /-- A slight generalisation of the **Marica-Schönheim Inequality**. -/ lemma Finset.le_card_diffs_mul_card_diffs (s t : Finset α) : - s.card * t.card ≤ (s \\ t).card * (t \\ s).card := by + #s * #t ≤ #(s \\ t) * #(t \\ s) := by have : ∀ s t : Finset α, (s \\ t).map ⟨_, liftLatticeHom_injective⟩ = s.map ⟨_, liftLatticeHom_injective⟩ \\ t.map ⟨_, liftLatticeHom_injective⟩ := by rintro s t @@ -364,6 +364,6 @@ lemma Finset.le_card_diffs_mul_card_diffs (s t : Finset α) : (t.map ⟨_, liftLatticeHom_injective⟩)ᶜˢ /-- The **Marica-Schönheim Inequality**. -/ -lemma Finset.card_le_card_diffs (s : Finset α) : s.card ≤ (s \\ s).card := +lemma Finset.card_le_card_diffs (s : Finset α) : #s ≤ #(s \\ s) := le_of_pow_le_pow_left two_ne_zero (zero_le _) <| by simpa [← sq] using s.le_card_diffs_mul_card_diffs s diff --git a/Mathlib/Combinatorics/SetFamily/HarrisKleitman.lean b/Mathlib/Combinatorics/SetFamily/HarrisKleitman.lean index 2a5e38dc5b5f6..49205ae8ec3d0 100644 --- a/Mathlib/Combinatorics/SetFamily/HarrisKleitman.lean +++ b/Mathlib/Combinatorics/SetFamily/HarrisKleitman.lean @@ -11,8 +11,8 @@ import Mathlib.Data.Fintype.Powerset /-! # Harris-Kleitman inequality -This file proves the Harris-Kleitman inequality. This relates `𝒜.card * ℬ.card` and -`2 ^ card α * (𝒜 ∩ ℬ).card` where `𝒜` and `ℬ` are upward- or downcard-closed finite families of +This file proves the Harris-Kleitman inequality. This relates `#𝒜 * #ℬ` and +`2 ^ card α * #(𝒜 ∩ ℬ)` where `𝒜` and `ℬ` are upward- or downcard-closed finite families of finsets. This can be interpreted as saying that any two lower sets (resp. any two upper sets) correlate in the uniform measure. @@ -49,7 +49,7 @@ theorem IsLowerSet.memberSubfamily_subset_nonMemberSubfamily (h : IsLowerSet ( /-- **Harris-Kleitman inequality**: Any two lower sets of finsets correlate. -/ theorem IsLowerSet.le_card_inter_finset' (h𝒜 : IsLowerSet (𝒜 : Set (Finset α))) (hℬ : IsLowerSet (ℬ : Set (Finset α))) (h𝒜s : ∀ t ∈ 𝒜, t ⊆ s) (hℬs : ∀ t ∈ ℬ, t ⊆ s) : - 𝒜.card * ℬ.card ≤ 2 ^ s.card * (𝒜 ∩ ℬ).card := by + #𝒜 * #ℬ ≤ 2 ^ #s * #(𝒜 ∩ ℬ) := by induction' s using Finset.induction with a s hs ih generalizing 𝒜 ℬ · simp_rw [subset_empty, ← subset_singleton_iff', subset_singleton_iff] at h𝒜s hℬs obtain rfl | rfl := h𝒜s @@ -89,13 +89,13 @@ variable [Fintype α] /-- **Harris-Kleitman inequality**: Any two lower sets of finsets correlate. -/ theorem IsLowerSet.le_card_inter_finset (h𝒜 : IsLowerSet (𝒜 : Set (Finset α))) - (hℬ : IsLowerSet (ℬ : Set (Finset α))) : 𝒜.card * ℬ.card ≤ 2 ^ Fintype.card α * (𝒜 ∩ ℬ).card := + (hℬ : IsLowerSet (ℬ : Set (Finset α))) : #𝒜 * #ℬ ≤ 2 ^ Fintype.card α * #(𝒜 ∩ ℬ) := h𝒜.le_card_inter_finset' hℬ (fun _ _ => subset_univ _) fun _ _ => subset_univ _ /-- **Harris-Kleitman inequality**: Upper sets and lower sets of finsets anticorrelate. -/ theorem IsUpperSet.card_inter_le_finset (h𝒜 : IsUpperSet (𝒜 : Set (Finset α))) (hℬ : IsLowerSet (ℬ : Set (Finset α))) : - 2 ^ Fintype.card α * (𝒜 ∩ ℬ).card ≤ 𝒜.card * ℬ.card := by + 2 ^ Fintype.card α * #(𝒜 ∩ ℬ) ≤ #𝒜 * #ℬ := by rw [← isLowerSet_compl, ← coe_compl] at h𝒜 have := h𝒜.le_card_inter_finset hℬ rwa [card_compl, Fintype.card_finset, tsub_mul, tsub_le_iff_tsub_le, ← mul_tsub, ← @@ -105,14 +105,14 @@ theorem IsUpperSet.card_inter_le_finset (h𝒜 : IsUpperSet (𝒜 : Set (Finset /-- **Harris-Kleitman inequality**: Lower sets and upper sets of finsets anticorrelate. -/ theorem IsLowerSet.card_inter_le_finset (h𝒜 : IsLowerSet (𝒜 : Set (Finset α))) (hℬ : IsUpperSet (ℬ : Set (Finset α))) : - 2 ^ Fintype.card α * (𝒜 ∩ ℬ).card ≤ 𝒜.card * ℬ.card := by - rw [inter_comm, mul_comm 𝒜.card] + 2 ^ Fintype.card α * #(𝒜 ∩ ℬ) ≤ #𝒜 * #ℬ := by + rw [inter_comm, mul_comm #𝒜] exact hℬ.card_inter_le_finset h𝒜 /-- **Harris-Kleitman inequality**: Any two upper sets of finsets correlate. -/ theorem IsUpperSet.le_card_inter_finset (h𝒜 : IsUpperSet (𝒜 : Set (Finset α))) (hℬ : IsUpperSet (ℬ : Set (Finset α))) : - 𝒜.card * ℬ.card ≤ 2 ^ Fintype.card α * (𝒜 ∩ ℬ).card := by + #𝒜 * #ℬ ≤ 2 ^ Fintype.card α * #(𝒜 ∩ ℬ) := by rw [← isLowerSet_compl, ← coe_compl] at h𝒜 have := h𝒜.card_inter_le_finset hℬ rwa [card_compl, Fintype.card_finset, tsub_mul, le_tsub_iff_le_tsub, ← mul_tsub, ← diff --git a/Mathlib/Combinatorics/SetFamily/Intersecting.lean b/Mathlib/Combinatorics/SetFamily/Intersecting.lean index 11d2839e67959..456e133bdb79d 100644 --- a/Mathlib/Combinatorics/SetFamily/Intersecting.lean +++ b/Mathlib/Combinatorics/SetFamily/Intersecting.lean @@ -141,7 +141,7 @@ theorem Intersecting.disjoint_map_compl {s : Finset α} (hs : (s : Set α).Inter exact hs.not_compl_mem hx' hx theorem Intersecting.card_le [Fintype α] {s : Finset α} (hs : (s : Set α).Intersecting) : - 2 * s.card ≤ Fintype.card α := by + 2 * #s ≤ Fintype.card α := by classical refine (s.disjUnion _ hs.disjoint_map_compl).card_le_univ.trans_eq' ?_ rw [Nat.two_mul, card_disjUnion, card_map] @@ -150,7 +150,7 @@ variable [Nontrivial α] [Fintype α] {s : Finset α} -- Note, this lemma is false when `α` has exactly one element and boring when `α` is empty. theorem Intersecting.is_max_iff_card_eq (hs : (s : Set α).Intersecting) : - (∀ t : Finset α, (t : Set α).Intersecting → s ⊆ t → s = t) ↔ 2 * s.card = Fintype.card α := by + (∀ t : Finset α, (t : Set α).Intersecting → s ⊆ t → s = t) ↔ 2 * #s = Fintype.card α := by classical refine ⟨fun h ↦ ?_, fun h t ht hst ↦ Finset.eq_of_subset_of_card_le hst <| Nat.le_of_mul_le_mul_left (ht.card_le.trans_eq h.symm) Nat.two_pos⟩ @@ -171,7 +171,7 @@ theorem Intersecting.is_max_iff_card_eq (hs : (s : Set α).Intersecting) : exact Finset.singleton_ne_empty _ (this <| Finset.empty_subset _).symm theorem Intersecting.exists_card_eq (hs : (s : Set α).Intersecting) : - ∃ t, s ⊆ t ∧ 2 * t.card = Fintype.card α ∧ (t : Set α).Intersecting := by + ∃ t, s ⊆ t ∧ 2 * #t = Fintype.card α ∧ (t : Set α).Intersecting := by have := hs.card_le rw [mul_comm, ← Nat.le_div_iff_mul_le' Nat.two_pos] at this revert hs diff --git a/Mathlib/Combinatorics/SetFamily/Kleitman.lean b/Mathlib/Combinatorics/SetFamily/Kleitman.lean index e534f30f51b45..dfb8f4f8241c7 100644 --- a/Mathlib/Combinatorics/SetFamily/Kleitman.lean +++ b/Mathlib/Combinatorics/SetFamily/Kleitman.lean @@ -34,11 +34,11 @@ variable {ι α : Type*} [Fintype α] [DecidableEq α] [Nonempty α] each further intersecting family takes at most half of the sets that are in no previous family. -/ theorem Finset.card_biUnion_le_of_intersecting (s : Finset ι) (f : ι → Finset (Finset α)) (hf : ∀ i ∈ s, (f i : Set (Finset α)).Intersecting) : - (s.biUnion f).card ≤ 2 ^ Fintype.card α - 2 ^ (Fintype.card α - s.card) := by + #(s.biUnion f) ≤ 2 ^ Fintype.card α - 2 ^ (Fintype.card α - #s) := by have : DecidableEq ι := by classical infer_instance - obtain hs | hs := le_total (Fintype.card α) s.card + obtain hs | hs := le_total (Fintype.card α) #s · rw [tsub_eq_zero_of_le hs, pow_zero] refine (card_le_card <| biUnion_subset.2 fun i hi a ha ↦ mem_compl.2 <| not_mem_singleton.2 <| (hf _ hi).ne_bot ha).trans_eq ?_ @@ -47,7 +47,7 @@ theorem Finset.card_biUnion_le_of_intersecting (s : Finset ι) (f : ι → Finse · simp set f' : ι → Finset (Finset α) := fun j ↦ if hj : j ∈ cons i s hi then (hf j hj).exists_card_eq.choose else ∅ - have hf₁ : ∀ j, j ∈ cons i s hi → f j ⊆ f' j ∧ 2 * (f' j).card = + have hf₁ : ∀ j, j ∈ cons i s hi → f j ⊆ f' j ∧ 2 * #(f' j) = 2 ^ Fintype.card α ∧ (f' j : Set (Finset α)).Intersecting := by rintro j hj simp_rw [f', dif_pos hj, ← Fintype.card_finset] diff --git a/Mathlib/Combinatorics/SetFamily/KruskalKatona.lean b/Mathlib/Combinatorics/SetFamily/KruskalKatona.lean index 23daec75ae3d2..f8957f5d2c015 100644 --- a/Mathlib/Combinatorics/SetFamily/KruskalKatona.lean +++ b/Mathlib/Combinatorics/SetFamily/KruskalKatona.lean @@ -78,8 +78,8 @@ lemma shadow_initSeg [Fintype α] (hs : s.Nonempty) : -- Assume first t < s - min s, and take k as the colex witness for this have hjk : j ≤ k := min'_le _ _ (mem_compl.2 ‹k ∉ t›) have : j ∉ t := mem_compl.1 (min'_mem _ _) - have hcard : card s = card (insert j t) := by - rw [card_insert_of_not_mem ‹j ∉ t›, ← ‹_ = card t›, card_erase_add_one (min'_mem _ _)] + have hcard : #s = #(insert j t) := by + rw [card_insert_of_not_mem ‹j ∉ t›, ← ‹_ = #t›, card_erase_add_one (min'_mem _ _)] refine ⟨j, ‹_›, hcard, ?_⟩ -- Cases on j < k or j = k obtain hjk | r₁ := hjk.lt_or_eq @@ -139,7 +139,7 @@ lemma toColex_compress_lt_toColex {hU : U.Nonempty} {hV : V.Nonempty} (h : max' /-- These are the compressions which we will apply to decrease the "measure" of a family of sets.-/ private def UsefulCompression (U V : Finset α) : Prop := - Disjoint U V ∧ U.card = V.card ∧ ∃ (HU : U.Nonempty) (HV : V.Nonempty), max' U HU < max' V HV + Disjoint U V ∧ #U = #V ∧ ∃ (HU : U.Nonempty) (HV : V.Nonempty), max' U HU < max' V HV private instance UsefulCompression.instDecidableRel : @DecidableRel (Finset α) UsefulCompression := fun _ _ ↦ inferInstanceAs (Decidable (_ ∧ _)) @@ -148,8 +148,8 @@ private instance UsefulCompression.instDecidableRel : @DecidableRel (Finset α) shadow. In particular, 'good' means it's useful, and every smaller compression won't make a difference. -/ private lemma compression_improved (𝒜 : Finset (Finset α)) (h₁ : UsefulCompression U V) - (h₂ : ∀ ⦃U₁ V₁⦄, UsefulCompression U₁ V₁ → U₁.card < U.card → IsCompressed U₁ V₁ 𝒜) : - (∂ (𝓒 U V 𝒜)).card ≤ (∂ 𝒜).card := by + (h₂ : ∀ ⦃U₁ V₁⦄, UsefulCompression U₁ V₁ → #U₁ < #U → IsCompressed U₁ V₁ 𝒜) : + #(∂ (𝓒 U V 𝒜)) ≤ #(∂ 𝒜) := by obtain ⟨UVd, same_size, hU, hV, max_lt⟩ := h₁ refine card_shadow_compression_le _ _ fun x Hx ↦ ⟨min' V hV, min'_mem _ _, ?_⟩ obtain hU' | hU' := eq_or_lt_of_le (succ_le_iff.2 hU.card_pos) @@ -174,7 +174,7 @@ lemma isInitSeg_of_compressed {ℬ : Finset (Finset α)} {r : ℕ} (h₁ : (ℬ rintro A B hA ⟨hBA, sizeA⟩ by_contra hB have hAB : A ≠ B := ne_of_mem_of_not_mem hA hB - have hAB' : A.card = B.card := (h₁ hA).trans sizeA.symm + have hAB' : #A = #B := (h₁ hA).trans sizeA.symm have hU : (A \ B).Nonempty := sdiff_nonempty.2 fun h ↦ hAB <| eq_of_subset_of_card_le h hAB'.ge have hV : (B \ A).Nonempty := sdiff_nonempty.2 fun h ↦ hAB.symm <| eq_of_subset_of_card_le h hAB'.le @@ -203,14 +203,14 @@ private lemma familyMeasure_compression_lt_familyMeasure {U V : Finset (Fin n)} {hV : V.Nonempty} (h : max' U hU < max' V hV) {𝒜 : Finset (Finset (Fin n))} (a : 𝓒 U V 𝒜 ≠ 𝒜) : familyMeasure (𝓒 U V 𝒜) < familyMeasure 𝒜 := by rw [compression] at a ⊢ - have q : ∀ Q ∈ 𝒜.filter fun A ↦ compress U V A ∉ 𝒜, compress U V Q ≠ Q := by + have q : ∀ Q ∈ {A ∈ 𝒜 | compress U V A ∉ 𝒜}, compress U V Q ≠ Q := by simp_rw [mem_filter] intro Q hQ h rw [h] at hQ exact hQ.2 hQ.1 - have uA : (𝒜.filter fun A => compress U V A ∈ 𝒜) ∪ 𝒜.filter (fun A ↦ compress U V A ∉ 𝒜) = 𝒜 := + have uA : {A ∈ 𝒜 | compress U V A ∈ 𝒜} ∪ {A ∈ 𝒜 | compress U V A ∉ 𝒜} = 𝒜 := filter_union_filter_neg_eq _ _ - have ne₂ : (𝒜.filter fun A ↦ compress U V A ∉ 𝒜).Nonempty := by + have ne₂ : {A ∈ 𝒜 | compress U V A ∉ 𝒜}.Nonempty := by refine nonempty_iff_ne_empty.2 fun z ↦ a ?_ rw [filter_image, z, image_empty, union_empty] rwa [z, union_empty] at uA @@ -229,12 +229,12 @@ we can't any more, which gives a set family which is fully compressed and has th want. -/ private lemma kruskal_katona_helper {r : ℕ} (𝒜 : Finset (Finset (Fin n))) (h : (𝒜 : Set (Finset (Fin n))).Sized r) : - ∃ ℬ : Finset (Finset (Fin n)), (∂ ℬ).card ≤ (∂ 𝒜).card ∧ 𝒜.card = ℬ.card ∧ + ∃ ℬ : Finset (Finset (Fin n)), #(∂ ℬ) ≤ #(∂ 𝒜) ∧ #𝒜 = #ℬ ∧ (ℬ : Set (Finset (Fin n))).Sized r ∧ ∀ U V, UsefulCompression U V → IsCompressed U V ℬ := by classical -- Are there any compressions we can make now? set usable : Finset (Finset (Fin n) × Finset (Fin n)) := - univ.filter fun t ↦ UsefulCompression t.1 t.2 ∧ ¬ IsCompressed t.1 t.2 𝒜 + {t | UsefulCompression t.1 t.2 ∧ ¬ IsCompressed t.1 t.2 𝒜} obtain husable | husable := usable.eq_empty_or_nonempty -- No. Then where we are is the required set family. · refine ⟨𝒜, le_rfl, rfl, h, fun U V hUV ↦ ?_⟩ @@ -242,13 +242,13 @@ private lemma kruskal_katona_helper {r : ℕ} (𝒜 : Finset (Finset (Fin n))) by_contra h exact husable ⟨U, V⟩ <| mem_filter.2 ⟨mem_univ _, hUV, h⟩ -- Yes. Then apply the smallest compression, then keep going - obtain ⟨⟨U, V⟩, hUV, t⟩ := exists_min_image usable (fun t ↦ t.1.card) husable + obtain ⟨⟨U, V⟩, hUV, t⟩ := exists_min_image usable (fun t ↦ #t.1) husable rw [mem_filter] at hUV - have h₂ : ∀ U₁ V₁, UsefulCompression U₁ V₁ → U₁.card < U.card → IsCompressed U₁ V₁ 𝒜 := by + have h₂ : ∀ U₁ V₁, UsefulCompression U₁ V₁ → #U₁ < #U → IsCompressed U₁ V₁ 𝒜 := by rintro U₁ V₁ huseful hUcard by_contra h exact hUcard.not_le <| t ⟨U₁, V₁⟩ <| mem_filter.2 ⟨mem_univ _, huseful, h⟩ - have p1 : (∂ (𝓒 U V 𝒜)).card ≤ (∂ 𝒜).card := compression_improved _ hUV.2.1 h₂ + have p1 : #(∂ (𝓒 U V 𝒜)) ≤ #(∂ 𝒜) := compression_improved _ hUV.2.1 h₂ obtain ⟨-, hUV', hu, hv, hmax⟩ := hUV.2.1 have := familyMeasure_compression_lt_familyMeasure hmax hUV.2.2 obtain ⟨t, q1, q2, q3, q4⟩ := UV.kruskal_katona_helper (𝓒 U V 𝒜) (h.uvCompression hUV') @@ -266,8 +266,8 @@ variable {r k i : ℕ} {𝒜 𝒞 : Finset <| Finset <| Fin n} Given a set family `𝒜` consisting of `r`-sets, and `𝒞` an initial segment of the colex order of the same size, the shadow of `𝒞` is smaller than the shadow of `𝒜`. In particular, this gives that the minimum shadow size is achieved by initial segments of colex. -/ -theorem kruskal_katona (h𝒜r : (𝒜 : Set (Finset (Fin n))).Sized r) (h𝒞𝒜 : 𝒞.card ≤ 𝒜.card) - (h𝒞 : IsInitSeg 𝒞 r) : (∂ 𝒞).card ≤ (∂ 𝒜).card := by +theorem kruskal_katona (h𝒜r : (𝒜 : Set (Finset (Fin n))).Sized r) (h𝒞𝒜 : #𝒞 ≤ #𝒜) + (h𝒞 : IsInitSeg 𝒞 r) : #(∂ 𝒞) ≤ #(∂ 𝒜) := by -- WLOG `|𝒜| = |𝒞|` obtain ⟨𝒜', h𝒜, h𝒜𝒞⟩ := exists_subset_card_eq h𝒞𝒜 -- By `kruskal_katona_helper`, we find a fully compressed family `ℬ` of the same size as `𝒜` @@ -276,15 +276,15 @@ theorem kruskal_katona (h𝒜r : (𝒜 : Set (Finset (Fin n))).Sized r) (h𝒞 -- This means that `ℬ` is an initial segment of the same size as `𝒞`. Hence they are equal and -- we are done. suffices ℬ = 𝒞 by subst 𝒞; exact hℬ𝒜.trans (by gcongr) - have hcard : card ℬ = card 𝒞 := h𝒜ℬ.symm.trans h𝒜𝒞 + have hcard : #ℬ = #𝒞 := h𝒜ℬ.symm.trans h𝒜𝒞 obtain h𝒞ℬ | hℬ𝒞 := h𝒞.total (UV.isInitSeg_of_compressed hℬr hℬ) · exact (eq_of_subset_of_card_le h𝒞ℬ hcard.le).symm · exact eq_of_subset_of_card_le hℬ𝒞 hcard.ge /-- An iterated form of the Kruskal-Katona theorem. In particular, the minimum possible iterated shadow size is attained by initial segments. -/ -theorem iterated_kk (h₁ : (𝒜 : Set (Finset (Fin n))).Sized r) (h₂ : 𝒞.card ≤ 𝒜.card) - (h₃ : IsInitSeg 𝒞 r) : (∂^[k] 𝒞).card ≤ (∂^[k] 𝒜).card := by +theorem iterated_kk (h₁ : (𝒜 : Set (Finset (Fin n))).Sized r) (h₂ : #𝒞 ≤ #𝒜) (h₃ : IsInitSeg 𝒞 r) : + #(∂^[k] 𝒞) ≤ #(∂^[k] 𝒜) := by induction' k with _k ih generalizing r 𝒜 𝒞 · simpa · refine ih h₁.shadow (kruskal_katona h₁ h₂ h₃) ?_ diff --git a/Mathlib/Combinatorics/SetFamily/LYM.lean b/Mathlib/Combinatorics/SetFamily/LYM.lean index 6b5887d58e1ee..a203b51c8faf5 100644 --- a/Mathlib/Combinatorics/SetFamily/LYM.lean +++ b/Mathlib/Combinatorics/SetFamily/LYM.lean @@ -61,7 +61,7 @@ variable [DecidableEq α] [Fintype α] /-- The downward **local LYM inequality**, with cancelled denominators. `𝒜` takes up less of `α^(r)` (the finsets of card `r`) than `∂𝒜` takes up of `α^(r - 1)`. -/ theorem card_mul_le_card_shadow_mul (h𝒜 : (𝒜 : Set (Finset α)).Sized r) : - 𝒜.card * r ≤ (∂ 𝒜).card * (Fintype.card α - r + 1) := by + #𝒜 * r ≤ #(∂ 𝒜) * (Fintype.card α - r + 1) := by let i : DecidableRel ((· ⊆ ·) : Finset α → Finset α → Prop) := fun _ _ => Classical.dec _ refine card_mul_le_card_mul' (· ⊆ ·) (fun s hs => ?_) (fun s hs => ?_) · rw [← h𝒜 hs, ← card_image_of_injOn s.erase_injOn] @@ -87,8 +87,8 @@ theorem card_mul_le_card_shadow_mul (h𝒜 : (𝒜 : Set (Finset α)).Sized r) : /-- The downward **local LYM inequality**. `𝒜` takes up less of `α^(r)` (the finsets of card `r`) than `∂𝒜` takes up of `α^(r - 1)`. -/ theorem card_div_choose_le_card_shadow_div_choose (hr : r ≠ 0) - (h𝒜 : (𝒜 : Set (Finset α)).Sized r) : (𝒜.card : 𝕜) / (Fintype.card α).choose r - ≤ (∂ 𝒜).card / (Fintype.card α).choose (r - 1) := by + (h𝒜 : (𝒜 : Set (Finset α)).Sized r) : (#𝒜 : 𝕜) / (Fintype.card α).choose r + ≤ #(∂ 𝒜) / (Fintype.card α).choose (r - 1) := by obtain hr' | hr' := lt_or_le (Fintype.card α) r · rw [choose_eq_zero_of_lt hr', cast_zero, div_zero] exact div_nonneg (cast_nonneg _) (cast_nonneg _) @@ -122,7 +122,7 @@ def falling : Finset (Finset α) := variable {𝒜 k} {s : Finset α} -theorem mem_falling : s ∈ falling k 𝒜 ↔ (∃ t ∈ 𝒜, s ⊆ t) ∧ s.card = k := by +theorem mem_falling : s ∈ falling k 𝒜 ↔ (∃ t ∈ 𝒜, s ⊆ t) ∧ #s = k := by simp_rw [falling, mem_sup, mem_powersetCard] aesop @@ -169,7 +169,7 @@ theorem IsAntichain.disjoint_slice_shadow_falling {m n : ℕ} theorem le_card_falling_div_choose [Fintype α] (hk : k ≤ Fintype.card α) (h𝒜 : IsAntichain (· ⊆ ·) (𝒜 : Set (Finset α))) : (∑ r ∈ range (k + 1), - ((𝒜 # (Fintype.card α - r)).card : 𝕜) / (Fintype.card α).choose (Fintype.card α - r)) ≤ + (#(𝒜 # (Fintype.card α - r)) : 𝕜) / (Fintype.card α).choose (Fintype.card α - r)) ≤ (falling (Fintype.card α - k) 𝒜).card / (Fintype.card α).choose (Fintype.card α - k) := by induction' k with k ih · simp only [tsub_zero, cast_one, cast_le, sum_singleton, div_one, choose_self, range_one, @@ -194,7 +194,7 @@ variable {𝒜 : Finset (Finset α)} {s : Finset α} {k : ℕ} proportion of elements it takes from each layer is less than `1`. -/ theorem sum_card_slice_div_choose_le_one [Fintype α] (h𝒜 : IsAntichain (· ⊆ ·) (𝒜 : Set (Finset α))) : - (∑ r ∈ range (Fintype.card α + 1), ((𝒜 # r).card : 𝕜) / (Fintype.card α).choose r) ≤ 1 := by + (∑ r ∈ range (Fintype.card α + 1), (#(𝒜 # r) : 𝕜) / (Fintype.card α).choose r) ≤ 1 := by classical rw [← sum_flip] refine (le_card_falling_div_choose le_rfl h𝒜).trans ?_ @@ -213,10 +213,10 @@ end LYM maximal layer in `Finset α`. This precisely means that `Finset α` is a Sperner order. -/ theorem IsAntichain.sperner [Fintype α] {𝒜 : Finset (Finset α)} (h𝒜 : IsAntichain (· ⊆ ·) (𝒜 : Set (Finset α))) : - 𝒜.card ≤ (Fintype.card α).choose (Fintype.card α / 2) := by + #𝒜 ≤ (Fintype.card α).choose (Fintype.card α / 2) := by classical suffices (∑ r ∈ Iic (Fintype.card α), - ((𝒜 # r).card : ℚ) / (Fintype.card α).choose (Fintype.card α / 2)) ≤ 1 by + (#(𝒜 # r) : ℚ) / (Fintype.card α).choose (Fintype.card α / 2)) ≤ 1 by rw [← sum_div, ← Nat.cast_sum, div_le_one] at this · simp only [cast_le] at this rwa [sum_card_slice] at this diff --git a/Mathlib/Combinatorics/SetFamily/Shadow.lean b/Mathlib/Combinatorics/SetFamily/Shadow.lean index a3b293b819c4c..7e12901d9f2f6 100644 --- a/Mathlib/Combinatorics/SetFamily/Shadow.lean +++ b/Mathlib/Combinatorics/SetFamily/Shadow.lean @@ -97,7 +97,7 @@ theorem erase_mem_shadow (hs : s ∈ 𝒜) (ha : a ∈ s) : erase s a ∈ ∂ /-- `t ∈ ∂𝒜` iff `t` is exactly one element less than something from `𝒜`. See also `Finset.mem_shadow_iff_exists_mem_card_add_one`. -/ -lemma mem_shadow_iff_exists_sdiff : t ∈ ∂ 𝒜 ↔ ∃ s ∈ 𝒜, t ⊆ s ∧ (s \ t).card = 1 := by +lemma mem_shadow_iff_exists_sdiff : t ∈ ∂ 𝒜 ↔ ∃ s ∈ 𝒜, t ⊆ s ∧ #(s \ t) = 1 := by simp_rw [mem_shadow_iff, ← covBy_iff_card_sdiff_eq_one, covBy_iff_exists_erase] /-- `t` is in the shadow of `𝒜` iff we can add an element to it so that the resulting finset is in @@ -109,15 +109,14 @@ lemma mem_shadow_iff_insert_mem : t ∈ ∂ 𝒜 ↔ ∃ a ∉ t, insert a t ∈ /-- `s ∈ ∂ 𝒜` iff `s` is exactly one element less than something from `𝒜`. See also `Finset.mem_shadow_iff_exists_sdiff`. -/ -lemma mem_shadow_iff_exists_mem_card_add_one : - t ∈ ∂ 𝒜 ↔ ∃ s ∈ 𝒜, t ⊆ s ∧ s.card = t.card + 1 := by +lemma mem_shadow_iff_exists_mem_card_add_one : t ∈ ∂ 𝒜 ↔ ∃ s ∈ 𝒜, t ⊆ s ∧ #s = #t + 1 := by refine mem_shadow_iff_exists_sdiff.trans <| exists_congr fun t ↦ and_congr_right fun _ ↦ and_congr_right fun hst ↦ ?_ rw [card_sdiff hst, tsub_eq_iff_eq_add_of_le, add_comm] exact card_mono hst lemma mem_shadow_iterate_iff_exists_card : - t ∈ ∂^[k] 𝒜 ↔ ∃ u : Finset α, u.card = k ∧ Disjoint t u ∧ t ∪ u ∈ 𝒜 := by + t ∈ ∂^[k] 𝒜 ↔ ∃ u : Finset α, #u = k ∧ Disjoint t u ∧ t ∪ u ∈ 𝒜 := by induction' k with k ih generalizing t · simp set_option tactic.skipAssignedInstances false in @@ -127,7 +126,7 @@ lemma mem_shadow_iterate_iff_exists_card : /-- `t ∈ ∂^k 𝒜` iff `t` is exactly `k` elements less than something from `𝒜`. See also `Finset.mem_shadow_iff_exists_mem_card_add`. -/ -lemma mem_shadow_iterate_iff_exists_sdiff : t ∈ ∂^[k] 𝒜 ↔ ∃ s ∈ 𝒜, t ⊆ s ∧ (s \ t).card = k := by +lemma mem_shadow_iterate_iff_exists_sdiff : t ∈ ∂^[k] 𝒜 ↔ ∃ s ∈ 𝒜, t ⊆ s ∧ #(s \ t) = k := by rw [mem_shadow_iterate_iff_exists_card] constructor · rintro ⟨u, rfl, htu, hsuA⟩ @@ -140,7 +139,7 @@ lemma mem_shadow_iterate_iff_exists_sdiff : t ∈ ∂^[k] 𝒜 ↔ ∃ s ∈ See also `Finset.mem_shadow_iterate_iff_exists_sdiff`. -/ lemma mem_shadow_iterate_iff_exists_mem_card_add : - t ∈ ∂^[k] 𝒜 ↔ ∃ s ∈ 𝒜, t ⊆ s ∧ s.card = t.card + k := by + t ∈ ∂^[k] 𝒜 ↔ ∃ s ∈ 𝒜, t ⊆ s ∧ #s = #t + k := by refine mem_shadow_iterate_iff_exists_sdiff.trans <| exists_congr fun t ↦ and_congr_right fun _ ↦ and_congr_right fun hst ↦ ?_ rw [card_sdiff hst, tsub_eq_iff_eq_add_of_le, add_comm] @@ -209,7 +208,7 @@ theorem insert_mem_upShadow (hs : s ∈ 𝒜) (ha : a ∉ s) : insert a s ∈ /-- `t` is in the upper shadow of `𝒜` iff `t` is exactly one element more than something from `𝒜`. See also `Finset.mem_upShadow_iff_exists_mem_card_add_one`. -/ -lemma mem_upShadow_iff_exists_sdiff : t ∈ ∂⁺ 𝒜 ↔ ∃ s ∈ 𝒜, s ⊆ t ∧ (t \ s).card = 1 := by +lemma mem_upShadow_iff_exists_sdiff : t ∈ ∂⁺ 𝒜 ↔ ∃ s ∈ 𝒜, s ⊆ t ∧ #(t \ s) = 1 := by simp_rw [mem_upShadow_iff, ← covBy_iff_card_sdiff_eq_one, covBy_iff_exists_insert] /-- `t` is in the upper shadow of `𝒜` iff we can remove an element from it so that the resulting @@ -222,14 +221,14 @@ lemma mem_upShadow_iff_erase_mem : t ∈ ∂⁺ 𝒜 ↔ ∃ a, a ∈ t ∧ eras See also `Finset.mem_upShadow_iff_exists_sdiff`. -/ lemma mem_upShadow_iff_exists_mem_card_add_one : - t ∈ ∂⁺ 𝒜 ↔ ∃ s ∈ 𝒜, s ⊆ t ∧ t.card = s.card + 1 := by + t ∈ ∂⁺ 𝒜 ↔ ∃ s ∈ 𝒜, s ⊆ t ∧ #t = #s + 1 := by refine mem_upShadow_iff_exists_sdiff.trans <| exists_congr fun t ↦ and_congr_right fun _ ↦ and_congr_right fun hst ↦ ?_ rw [card_sdiff hst, tsub_eq_iff_eq_add_of_le, add_comm] exact card_mono hst lemma mem_upShadow_iterate_iff_exists_card : - t ∈ ∂⁺^[k] 𝒜 ↔ ∃ u : Finset α, u.card = k ∧ u ⊆ t ∧ t \ u ∈ 𝒜 := by + t ∈ ∂⁺^[k] 𝒜 ↔ ∃ u : Finset α, #u = k ∧ u ⊆ t ∧ t \ u ∈ 𝒜 := by induction' k with k ih generalizing t · simp simp only [mem_upShadow_iff_erase_mem, ih, Function.iterate_succ_apply', card_eq_succ, @@ -244,8 +243,7 @@ lemma mem_upShadow_iterate_iff_exists_card : /-- `t` is in the upper shadow of `𝒜` iff `t` is exactly `k` elements less than something from `𝒜`. See also `Finset.mem_upShadow_iff_exists_mem_card_add`. -/ -lemma mem_upShadow_iterate_iff_exists_sdiff : - t ∈ ∂⁺^[k] 𝒜 ↔ ∃ s ∈ 𝒜, s ⊆ t ∧ (t \ s).card = k := by +lemma mem_upShadow_iterate_iff_exists_sdiff : t ∈ ∂⁺^[k] 𝒜 ↔ ∃ s ∈ 𝒜, s ⊆ t ∧ #(t \ s) = k := by rw [mem_upShadow_iterate_iff_exists_card] constructor · rintro ⟨u, rfl, hut, htu⟩ @@ -257,7 +255,7 @@ lemma mem_upShadow_iterate_iff_exists_sdiff : See also `Finset.mem_upShadow_iterate_iff_exists_sdiff`. -/ lemma mem_upShadow_iterate_iff_exists_mem_card_add : - t ∈ ∂⁺^[k] 𝒜 ↔ ∃ s ∈ 𝒜, s ⊆ t ∧ t.card = s.card + k := by + t ∈ ∂⁺^[k] 𝒜 ↔ ∃ s ∈ 𝒜, s ⊆ t ∧ #t = #s + k := by refine mem_upShadow_iterate_iff_exists_sdiff.trans <| exists_congr fun t ↦ and_congr_right fun _ ↦ and_congr_right fun hst ↦ ?_ rw [card_sdiff hst, tsub_eq_iff_eq_add_of_le, add_comm] @@ -277,7 +275,7 @@ theorem exists_subset_of_mem_upShadow (hs : s ∈ ∂⁺ 𝒜) : ∃ t ∈ 𝒜, /-- `t ∈ ∂^k 𝒜` iff `t` is exactly `k` elements more than something in `𝒜`. -/ theorem mem_upShadow_iff_exists_mem_card_add : - s ∈ ∂⁺ ^[k] 𝒜 ↔ ∃ t ∈ 𝒜, t ⊆ s ∧ t.card + k = s.card := by + s ∈ ∂⁺ ^[k] 𝒜 ↔ ∃ t ∈ 𝒜, t ⊆ s ∧ #t + k = #s := by induction' k with k ih generalizing 𝒜 s · refine ⟨fun hs => ⟨s, hs, Subset.refl _, rfl⟩, ?_⟩ rintro ⟨t, ht, hst, hcard⟩ diff --git a/Mathlib/Combinatorics/SetFamily/Shatter.lean b/Mathlib/Combinatorics/SetFamily/Shatter.lean index 74f21f70fbc02..3410b51bce8cc 100644 --- a/Mathlib/Combinatorics/SetFamily/Shatter.lean +++ b/Mathlib/Combinatorics/SetFamily/Shatter.lean @@ -72,7 +72,8 @@ lemma univ_shatters [Fintype α] : univ.Shatters s := rw [shatters_iff, powerset_univ]; simp_rw [univ_inter, image_id'] /-- The set family of sets that are shattered by `𝒜`. -/ -def shatterer (𝒜 : Finset (Finset α)) : Finset (Finset α) := (𝒜.biUnion powerset).filter 𝒜.Shatters +def shatterer (𝒜 : Finset (Finset α)) : Finset (Finset α) := + {s ∈ 𝒜.biUnion powerset | 𝒜.Shatters s} @[simp] lemma mem_shatterer : s ∈ 𝒜.shatterer ↔ 𝒜.Shatters s := by refine mem_filter.trans <| and_iff_right_of_imp fun h ↦ ?_ @@ -106,15 +107,14 @@ private lemma aux (h : ∀ t ∈ 𝒜, a ∉ t) (ht : 𝒜.Shatters t) : a ∉ t obtain ⟨u, hu, htu⟩ := ht.exists_superset; exact not_mem_mono htu <| h u hu /-- Pajor's variant of the **Sauer-Shelah lemma**. -/ -lemma card_le_card_shatterer (𝒜 : Finset (Finset α)) : 𝒜.card ≤ 𝒜.shatterer.card := by +lemma card_le_card_shatterer (𝒜 : Finset (Finset α)) : #𝒜 ≤ #𝒜.shatterer := by refine memberFamily_induction_on 𝒜 ?_ ?_ ?_ · simp · rfl intros a 𝒜 ih₀ ih₁ set ℬ : Finset (Finset α) := ((memberSubfamily a 𝒜).shatterer ∩ (nonMemberSubfamily a 𝒜).shatterer).image (insert a) - have hℬ : - ℬ.card = ((memberSubfamily a 𝒜).shatterer ∩ (nonMemberSubfamily a 𝒜).shatterer).card := by + have hℬ : #ℬ = #((memberSubfamily a 𝒜).shatterer ∩ (nonMemberSubfamily a 𝒜).shatterer) := by refine card_image_of_injOn <| insert_erase_invOn.2.injOn.mono ?_ simp only [coe_inter, Set.subset_def, Set.mem_inter_iff, mem_coe, Set.mem_setOf_eq, and_imp, mem_shatterer] @@ -183,7 +183,7 @@ def vcDim (𝒜 : Finset (Finset α)) : ℕ := 𝒜.shatterer.sup card @[gcongr] lemma vcDim_mono (h𝒜ℬ : 𝒜 ⊆ ℬ) : 𝒜.vcDim ≤ ℬ.vcDim := by unfold vcDim; gcongr -lemma Shatters.card_le_vcDim (hs : 𝒜.Shatters s) : s.card ≤ 𝒜.vcDim := le_sup <| mem_shatterer.2 hs +lemma Shatters.card_le_vcDim (hs : 𝒜.Shatters s) : #s ≤ 𝒜.vcDim := le_sup <| mem_shatterer.2 hs /-- Down-compressing decreases the VC-dimension. -/ lemma vcDim_compress_le (a : α) (𝒜 : Finset (Finset α)) : (𝓓 a 𝒜).vcDim ≤ 𝒜.vcDim := @@ -191,9 +191,9 @@ lemma vcDim_compress_le (a : α) (𝒜 : Finset (Finset α)) : (𝓓 a 𝒜).vcD /-- The **Sauer-Shelah lemma**. -/ lemma card_shatterer_le_sum_vcDim [Fintype α] : - 𝒜.shatterer.card ≤ ∑ k ∈ Iic 𝒜.vcDim, (Fintype.card α).choose k := by + #𝒜.shatterer ≤ ∑ k ∈ Iic 𝒜.vcDim, (Fintype.card α).choose k := by simp_rw [← card_univ, ← card_powersetCard] - refine (card_le_card fun s hs ↦ mem_biUnion.2 ⟨card s, ?_⟩).trans card_biUnion_le + refine (card_le_card fun s hs ↦ mem_biUnion.2 ⟨#s, ?_⟩).trans card_biUnion_le exact ⟨mem_Iic.2 (mem_shatterer.1 hs).card_le_vcDim, mem_powersetCard_univ.2 rfl⟩ end Finset diff --git a/Mathlib/Combinatorics/SimpleGraph/Basic.lean b/Mathlib/Combinatorics/SimpleGraph/Basic.lean index 69de374b0e114..40be1f3971a03 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Basic.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Basic.lean @@ -709,7 +709,7 @@ theorem neighborSet_union_compl_neighborSet_eq (G : SimpleGraph V) (v : V) : theorem card_neighborSet_union_compl_neighborSet [Fintype V] (G : SimpleGraph V) (v : V) [Fintype (G.neighborSet v ∪ Gᶜ.neighborSet v : Set V)] : - (Set.toFinset (G.neighborSet v ∪ Gᶜ.neighborSet v)).card = Fintype.card V - 1 := by + #(G.neighborSet v ∪ Gᶜ.neighborSet v).toFinset = Fintype.card V - 1 := by classical simp_rw [neighborSet_union_compl_neighborSet_eq, Set.toFinset_compl, Finset.card_compl, Set.toFinset_card, Set.card_singleton] diff --git a/Mathlib/Combinatorics/SimpleGraph/Clique.lean b/Mathlib/Combinatorics/SimpleGraph/Clique.lean index 4023561674eb6..fc597f6af613d 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Clique.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Clique.lean @@ -140,8 +140,8 @@ theorem isClique_map_finset_iff_of_nontrivial (ht : t.Nontrivial) : simpa using hs.map (f := f) theorem isClique_map_finset_iff : - (G.map f).IsClique t ↔ t.card ≤ 1 ∨ ∃ (s : Finset α), G.IsClique s ∧ s.map f = t := by - obtain (ht | ht) := le_or_lt t.card 1 + (G.map f).IsClique t ↔ #t ≤ 1 ∨ ∃ (s : Finset α), G.IsClique s ∧ s.map f = t := by + obtain (ht | ht) := le_or_lt #t 1 · simp only [ht, true_or, iff_true] exact IsClique.of_subsingleton <| card_le_one.1 ht rw [isClique_map_finset_iff_of_nontrivial, ← not_lt] @@ -164,9 +164,9 @@ variable {n : ℕ} {s : Finset α} /-- An `n`-clique in a graph is a set of `n` vertices which are pairwise connected. -/ structure IsNClique (n : ℕ) (s : Finset α) : Prop where clique : G.IsClique s - card_eq : s.card = n + card_eq : #s = n -theorem isNClique_iff : G.IsNClique n s ↔ G.IsClique s ∧ s.card = n := +theorem isNClique_iff : G.IsNClique n s ↔ G.IsClique s ∧ #s = n := ⟨fun h ↦ ⟨h.1, h.2⟩, fun h ↦ ⟨h.1, h.2⟩⟩ instance [DecidableEq α] [DecidableRel G.Adj] {n : ℕ} {s : Finset α} : @@ -199,7 +199,7 @@ theorem isNClique_map_iff (hn : 1 < n) {t : Finset β} {f : α ↪ β} : simp [hs.card_eq, hs.clique] @[simp] -theorem isNClique_bot_iff : (⊥ : SimpleGraph α).IsNClique n s ↔ n ≤ 1 ∧ s.card = n := by +theorem isNClique_bot_iff : (⊥ : SimpleGraph α).IsNClique n s ↔ n ≤ 1 ∧ #s = n := by rw [isNClique_iff, isClique_bot_iff] refine and_congr_left ?_ rintro rfl @@ -286,7 +286,7 @@ noncomputable def topEmbeddingOfNotCliqueFree {n : ℕ} (h : ¬G.CliqueFree n) : (⊤ : SimpleGraph (Fin n)) ↪g G := by simp only [CliqueFree, isNClique_iff, isClique_iff_induce_eq, not_forall, Classical.not_not] at h obtain ⟨ha, hb⟩ := h.choose_spec - have : (⊤ : SimpleGraph (Fin h.choose.card)) ≃g (⊤ : SimpleGraph h.choose) := by + have : (⊤ : SimpleGraph (Fin #h.choose)) ≃g (⊤ : SimpleGraph h.choose) := by apply Iso.completeGraph simpa using (Fintype.equivFin h.choose).symm rw [← ha] at this @@ -446,7 +446,7 @@ theorem cliqueFreeOn_univ : G.CliqueFreeOn Set.univ n ↔ G.CliqueFree n := by protected theorem CliqueFree.cliqueFreeOn (hG : G.CliqueFree n) : G.CliqueFreeOn s n := fun _t _ ↦ hG _ -theorem cliqueFreeOn_of_card_lt {s : Finset α} (h : s.card < n) : G.CliqueFreeOn s n := +theorem cliqueFreeOn_of_card_lt {s : Finset α} (h : #s < n) : G.CliqueFreeOn s n := fun _t hts ht => h.not_le <| ht.2.symm.trans_le <| card_mono hts -- TODO: Restate using `SimpleGraph.IndepSet` once we have it @@ -548,8 +548,7 @@ section CliqueFinset variable [Fintype α] [DecidableEq α] [DecidableRel G.Adj] {n : ℕ} {a b c : α} {s : Finset α} /-- The `n`-cliques in a graph as a finset. -/ -def cliqueFinset (n : ℕ) : Finset (Finset α) := - univ.filter <| G.IsNClique n +def cliqueFinset (n : ℕ) : Finset (Finset α) := {s | G.IsNClique n s} variable {G} in @[simp] @@ -568,7 +567,7 @@ theorem cliqueFinset_eq_empty_iff : G.cliqueFinset n = ∅ ↔ G.CliqueFree n := protected alias ⟨_, CliqueFree.cliqueFinset⟩ := cliqueFinset_eq_empty_iff -theorem card_cliqueFinset_le : (G.cliqueFinset n).card ≤ (card α).choose n := by +theorem card_cliqueFinset_le : #(G.cliqueFinset n) ≤ (card α).choose n := by rw [← card_univ, ← card_powersetCard] refine card_mono fun s => ?_ simpa [mem_powersetCard_univ] using IsNClique.card_eq diff --git a/Mathlib/Combinatorics/SimpleGraph/Connectivity/WalkCounting.lean b/Mathlib/Combinatorics/SimpleGraph/Connectivity/WalkCounting.lean index 2585da2220351..d1c68ef3928c6 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Connectivity/WalkCounting.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Connectivity/WalkCounting.lean @@ -21,7 +21,7 @@ can also be useful as a recursive description of this set when `V` is finite. TODO: should this be extended further? -/ -open Function +open Finset Function universe u v w @@ -152,7 +152,7 @@ theorem set_walk_length_toFinset_eq (n : ℕ) (u v : V) : /- See `SimpleGraph.adjMatrix_pow_apply_eq_card_walk` for the cardinality in terms of the `n`th power of the adjacency matrix. -/ theorem card_set_walk_length_eq (u v : V) (n : ℕ) : - Fintype.card {p : G.Walk u v | p.length = n} = (G.finsetWalkLength n u v).card := + Fintype.card {p : G.Walk u v | p.length = n} = #(G.finsetWalkLength n u v) := Fintype.card_ofFinset (G.finsetWalkLength n u v) fun p => by rw [← Finset.mem_coe, coe_finsetWalkLength_eq] @@ -165,7 +165,7 @@ instance fintypeSubtypeWalkLengthLT (u v : V) (n : ℕ) : Fintype {p : G.Walk u instance fintypeSetPathLength (u v : V) (n : ℕ) : Fintype {p : G.Walk u v | p.IsPath ∧ p.length = n} := - Fintype.ofFinset ((G.finsetWalkLength n u v).filter Walk.IsPath) <| by + Fintype.ofFinset {w ∈ G.finsetWalkLength n u v | w.IsPath} <| by simp [mem_finsetWalkLength_iff, and_comm] instance fintypeSubtypePathLength (u v : V) (n : ℕ) : @@ -174,7 +174,7 @@ instance fintypeSubtypePathLength (u v : V) (n : ℕ) : instance fintypeSetPathLengthLT (u v : V) (n : ℕ) : Fintype {p : G.Walk u v | p.IsPath ∧ p.length < n} := - Fintype.ofFinset ((G.finsetWalkLengthLT n u v).filter Walk.IsPath) <| by + Fintype.ofFinset {w ∈ G.finsetWalkLengthLT n u v | w.IsPath} <| by simp [mem_finsetWalkLengthLT_iff, and_comm] instance fintypeSubtypePathLengthLT (u v : V) (n : ℕ) : @@ -210,7 +210,6 @@ instance : Decidable G.Connected := by rw [connected_iff, ← Finset.univ_nonempty_iff] infer_instance -open Finset in instance Path.instFintype {u v : V} : Fintype (G.Path u v) where elems := (univ (α := { p : G.Walk u v | p.IsPath ∧ p.length < Fintype.card V })).map ⟨fun p ↦ { val := p.val, property := p.prop.left }, diff --git a/Mathlib/Combinatorics/SimpleGraph/DegreeSum.lean b/Mathlib/Combinatorics/SimpleGraph/DegreeSum.lean index 4b09938bdf173..5904dc6245e9b 100644 --- a/Mathlib/Combinatorics/SimpleGraph/DegreeSum.lean +++ b/Mathlib/Combinatorics/SimpleGraph/DegreeSum.lean @@ -49,7 +49,7 @@ section DegreeSum variable [Fintype V] [DecidableRel G.Adj] theorem dart_fst_fiber [DecidableEq V] (v : V) : - (univ.filter fun d : G.Dart => d.fst = v) = univ.image (G.dartOfNeighborSet v) := by + ({d : G.Dart | d.fst = v} : Finset _) = univ.image (G.dartOfNeighborSet v) := by ext d simp only [mem_image, true_and, mem_filter, SetCoe.exists, mem_univ, exists_prop_of_true] constructor @@ -59,7 +59,7 @@ theorem dart_fst_fiber [DecidableEq V] (v : V) : rfl theorem dart_fst_fiber_card_eq_degree [DecidableEq V] (v : V) : - (univ.filter fun d : G.Dart => d.fst = v).card = G.degree v := by + #{d : G.Dart | d.fst = v} = G.degree v := by simpa only [dart_fst_fiber, Finset.card_univ, card_neighborSet_eq_degree] using card_image_of_injective univ (G.dartOfNeighborSet_injective v) @@ -71,13 +71,13 @@ theorem dart_card_eq_sum_degrees : Fintype.card G.Dart = ∑ v, G.degree v := by variable {G} theorem Dart.edge_fiber [DecidableEq V] (d : G.Dart) : - (univ.filter fun d' : G.Dart => d'.edge = d.edge) = {d, d.symm} := + ({d' : G.Dart | d'.edge = d.edge} : Finset _) = {d, d.symm} := Finset.ext fun d' => by simpa using dart_edge_eq_iff d' d variable (G) theorem dart_edge_fiber_card [DecidableEq V] (e : Sym2 V) (h : e ∈ G.edgeSet) : - (univ.filter fun d : G.Dart => d.edge = e).card = 2 := by + #{d : G.Dart | d.edge = e} = 2 := by induction' e with v w let d : G.Dart := ⟨(v, w), h⟩ convert congr_arg card d.edge_fiber @@ -85,7 +85,7 @@ theorem dart_edge_fiber_card [DecidableEq V] (e : Sym2 V) (h : e ∈ G.edgeSet) rw [mem_singleton] exact d.symm_ne.symm -theorem dart_card_eq_twice_card_edges : Fintype.card G.Dart = 2 * G.edgeFinset.card := by +theorem dart_card_eq_twice_card_edges : Fintype.card G.Dart = 2 * #G.edgeFinset := by classical rw [← card_univ] rw [@card_eq_sum_card_fiberwise _ _ _ Dart.edge _ G.edgeFinset fun d _h => @@ -97,11 +97,10 @@ theorem dart_card_eq_twice_card_edges : Fintype.card G.Dart = 2 * G.edgeFinset.c /-- The degree-sum formula. This is also known as the handshaking lemma, which might more specifically refer to `SimpleGraph.even_card_odd_degree_vertices`. -/ -theorem sum_degrees_eq_twice_card_edges : ∑ v, G.degree v = 2 * G.edgeFinset.card := +theorem sum_degrees_eq_twice_card_edges : ∑ v, G.degree v = 2 * #G.edgeFinset := G.dart_card_eq_sum_degrees.symm.trans G.dart_card_eq_twice_card_edges -lemma two_mul_card_edgeFinset : - 2 * G.edgeFinset.card = (univ.filter fun (x, y) ↦ G.Adj x y).card := by +lemma two_mul_card_edgeFinset : 2 * #G.edgeFinset = #(univ.filter fun (x, y) ↦ G.Adj x y) := by rw [← dart_card_eq_twice_card_edges, ← card_univ] refine card_bij' (fun d _ ↦ (d.fst, d.snd)) (fun xy h ↦ ⟨xy, (mem_filter.1 h).2⟩) ?_ ?_ ?_ ?_ <;> simp @@ -110,7 +109,7 @@ end DegreeSum /-- The handshaking lemma. See also `SimpleGraph.sum_degrees_eq_twice_card_edges`. -/ theorem even_card_odd_degree_vertices [Fintype V] [DecidableRel G.Adj] : - Even (univ.filter fun v => Odd (G.degree v)).card := by + Even #{v | Odd (G.degree v)} := by classical have h := congr_arg (fun n => ↑n : ℕ → ZMod 2) G.sum_degrees_eq_twice_card_edges simp only [ZMod.natCast_self, zero_mul, Nat.cast_mul] at h @@ -126,10 +125,10 @@ theorem even_card_odd_degree_vertices [Fintype V] [DecidableRel G.Adj] : trivial theorem odd_card_odd_degree_vertices_ne [Fintype V] [DecidableEq V] [DecidableRel G.Adj] (v : V) - (h : Odd (G.degree v)) : Odd (univ.filter fun w => w ≠ v ∧ Odd (G.degree w)).card := by + (h : Odd (G.degree v)) : Odd #{w | w ≠ v ∧ Odd (G.degree w)} := by rcases G.even_card_odd_degree_vertices with ⟨k, hg⟩ have hk : 0 < k := by - have hh : (filter (fun v : V => Odd (G.degree v)) univ).Nonempty := by + have hh : Finset.Nonempty {v : V | Odd (G.degree v)} := by use v simp only [true_and, mem_filter, mem_univ] exact h @@ -150,7 +149,7 @@ theorem exists_ne_odd_degree_of_exists_odd_degree [Fintype V] [DecidableRel G.Ad (h : Odd (G.degree v)) : ∃ w : V, w ≠ v ∧ Odd (G.degree w) := by haveI := Classical.decEq V rcases G.odd_card_odd_degree_vertices_ne v h with ⟨k, hg⟩ - have hg' : (filter (fun w : V => w ≠ v ∧ Odd (G.degree w)) univ).card > 0 := by + have hg' : 0 < #{w | w ≠ v ∧ Odd (G.degree w)} := by rw [hg] apply Nat.succ_pos rcases card_pos.mp hg' with ⟨w, hw⟩ diff --git a/Mathlib/Combinatorics/SimpleGraph/Density.lean b/Mathlib/Combinatorics/SimpleGraph/Density.lean index f71bb81713dbb..6e75fd4155275 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Density.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Density.lean @@ -41,12 +41,10 @@ variable [LinearOrderedField 𝕜] (r : α → β → Prop) [∀ a, DecidablePre {t t₁ t₂ : Finset β} {a : α} {b : β} {δ : 𝕜} /-- Finset of edges of a relation between two finsets of vertices. -/ -def interedges (s : Finset α) (t : Finset β) : Finset (α × β) := - (s ×ˢ t).filter fun e ↦ r e.1 e.2 +def interedges (s : Finset α) (t : Finset β) : Finset (α × β) := {e ∈ s ×ˢ t | r e.1 e.2} /-- Edge density of a relation between two finsets of vertices. -/ -def edgeDensity (s : Finset α) (t : Finset β) : ℚ := - (interedges r s t).card / (s.card * t.card) +def edgeDensity (s : Finset α) (t : Finset β) : ℚ := #(interedges r s t) / (#s * #t) variable {r} @@ -68,7 +66,7 @@ theorem interedges_mono (hs : s₂ ⊆ s₁) (ht : t₂ ⊆ t₁) : interedges r variable (r) theorem card_interedges_add_card_interedges_compl (s : Finset α) (t : Finset β) : - (interedges r s t).card + (interedges (fun x y ↦ ¬r x y) s t).card = s.card * t.card := by + #(interedges r s t) + #(interedges (fun x y ↦ ¬r x y) s t) = #s * #t := by classical rw [← card_product, interedges, interedges, ← card_union_of_disjoint, filter_union_filter_neg_eq] exact disjoint_filter.2 fun _ _ ↦ Classical.not_not.2 @@ -92,7 +90,7 @@ section DecidableEq variable [DecidableEq α] [DecidableEq β] lemma interedges_eq_biUnion : - interedges r s t = s.biUnion (fun x ↦ (t.filter (r x)).map ⟨(x, ·), Prod.mk.inj_left x⟩) := by + interedges r s t = s.biUnion fun x ↦ {y ∈ t | r x y}.map ⟨(x, ·), Prod.mk.inj_left x⟩ := by ext ⟨x, y⟩; simp [mem_interedges_iff] theorem interedges_biUnion_left (s : Finset ι) (t : Finset β) (f : ι → Finset α) : @@ -115,7 +113,7 @@ theorem interedges_biUnion (s : Finset ι) (t : Finset κ) (f : ι → Finset α end DecidableEq theorem card_interedges_le_mul (s : Finset α) (t : Finset β) : - (interedges r s t).card ≤ s.card * t.card := + #(interedges r s t) ≤ #s * #t := (card_filter_le _ _).trans (card_product _ _).le theorem edgeDensity_nonneg (s : Finset α) (t : Finset β) : 0 ≤ edgeDensity r s t := by @@ -141,14 +139,14 @@ theorem edgeDensity_empty_right (s : Finset α) : edgeDensity r s ∅ = 0 := by rw [edgeDensity, Finset.card_empty, Nat.cast_zero, mul_zero, div_zero] theorem card_interedges_finpartition_left [DecidableEq α] (P : Finpartition s) (t : Finset β) : - (interedges r s t).card = ∑ a ∈ P.parts, (interedges r a t).card := by + #(interedges r s t) = ∑ a ∈ P.parts, #(interedges r a t) := by classical simp_rw [← P.biUnion_parts, interedges_biUnion_left, id] rw [card_biUnion] exact fun x hx y hy h ↦ interedges_disjoint_left r (P.disjoint hx hy h) _ theorem card_interedges_finpartition_right [DecidableEq β] (s : Finset α) (P : Finpartition t) : - (interedges r s t).card = ∑ b ∈ P.parts, (interedges r s b).card := by + #(interedges r s t) = ∑ b ∈ P.parts, #(interedges r s b) := by classical simp_rw [← P.biUnion_parts, interedges_biUnion_right, id] rw [card_biUnion] @@ -156,22 +154,22 @@ theorem card_interedges_finpartition_right [DecidableEq β] (s : Finset α) (P : theorem card_interedges_finpartition [DecidableEq α] [DecidableEq β] (P : Finpartition s) (Q : Finpartition t) : - (interedges r s t).card = ∑ ab ∈ P.parts ×ˢ Q.parts, (interedges r ab.1 ab.2).card := by + #(interedges r s t) = ∑ ab ∈ P.parts ×ˢ Q.parts, #(interedges r ab.1 ab.2) := by rw [card_interedges_finpartition_left _ P, sum_product] congr; ext rw [card_interedges_finpartition_right] theorem mul_edgeDensity_le_edgeDensity (hs : s₂ ⊆ s₁) (ht : t₂ ⊆ t₁) (hs₂ : s₂.Nonempty) (ht₂ : t₂.Nonempty) : - (s₂.card : ℚ) / s₁.card * (t₂.card / t₁.card) * edgeDensity r s₂ t₂ ≤ edgeDensity r s₁ t₁ := by - have hst : (s₂.card : ℚ) * t₂.card ≠ 0 := by simp [hs₂.ne_empty, ht₂.ne_empty] + (#s₂ : ℚ) / #s₁ * (#t₂ / #t₁) * edgeDensity r s₂ t₂ ≤ edgeDensity r s₁ t₁ := by + have hst : (#s₂ : ℚ) * #t₂ ≠ 0 := by simp [hs₂.ne_empty, ht₂.ne_empty] rw [edgeDensity, edgeDensity, div_mul_div_comm, mul_comm, div_mul_div_cancel₀ hst] gcongr exact interedges_mono hs ht theorem edgeDensity_sub_edgeDensity_le_one_sub_mul (hs : s₂ ⊆ s₁) (ht : t₂ ⊆ t₁) (hs₂ : s₂.Nonempty) (ht₂ : t₂.Nonempty) : - edgeDensity r s₂ t₂ - edgeDensity r s₁ t₁ ≤ 1 - s₂.card / s₁.card * (t₂.card / t₁.card) := by + edgeDensity r s₂ t₂ - edgeDensity r s₁ t₁ ≤ 1 - #s₂ / #s₁ * (#t₂ / #t₁) := by 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] @@ -182,7 +180,7 @@ theorem edgeDensity_sub_edgeDensity_le_one_sub_mul (hs : s₂ ⊆ s₁) (ht : t theorem abs_edgeDensity_sub_edgeDensity_le_one_sub_mul (hs : s₂ ⊆ s₁) (ht : t₂ ⊆ t₁) (hs₂ : s₂.Nonempty) (ht₂ : t₂.Nonempty) : - |edgeDensity r s₂ t₂ - edgeDensity r s₁ t₁| ≤ 1 - s₂.card / s₁.card * (t₂.card / t₁.card) := by + |edgeDensity r s₂ t₂ - edgeDensity r s₁ t₁| ≤ 1 - #s₂ / #s₁ * (#t₂ / #t₁) := by refine abs_sub_le_iff.2 ⟨edgeDensity_sub_edgeDensity_le_one_sub_mul r hs ht hs₂ ht₂, ?_⟩ rw [← add_sub_cancel_right (edgeDensity r s₁ t₁) (edgeDensity (fun x y ↦ ¬r x y) s₁ t₁), ← add_sub_cancel_right (edgeDensity r s₂ t₂) (edgeDensity (fun x y ↦ ¬r x y) s₂ t₂), @@ -191,8 +189,8 @@ theorem abs_edgeDensity_sub_edgeDensity_le_one_sub_mul (hs : s₂ ⊆ s₁) (ht exact edgeDensity_sub_edgeDensity_le_one_sub_mul _ hs ht hs₂ ht₂ theorem abs_edgeDensity_sub_edgeDensity_le_two_mul_sub_sq (hs : s₂ ⊆ s₁) (ht : t₂ ⊆ t₁) - (hδ₀ : 0 ≤ δ) (hδ₁ : δ < 1) (hs₂ : (1 - δ) * s₁.card ≤ s₂.card) - (ht₂ : (1 - δ) * t₁.card ≤ t₂.card) : + (hδ₀ : 0 ≤ δ) (hδ₁ : δ < 1) (hs₂ : (1 - δ) * #s₁ ≤ #s₂) + (ht₂ : (1 - δ) * #t₁ ≤ #t₂) : |(edgeDensity r s₂ t₂ : 𝕜) - edgeDensity r s₁ t₁| ≤ 2 * δ - δ ^ 2 := by have hδ' : 0 ≤ 2 * δ - δ ^ 2 := by rw [sub_nonneg, sq] @@ -222,7 +220,7 @@ theorem abs_edgeDensity_sub_edgeDensity_le_two_mul_sub_sq (hs : s₂ ⊆ s₁) ( /-- If `s₂ ⊆ s₁`, `t₂ ⊆ t₁` and they take up all but a `δ`-proportion, then the difference in edge densities is at most `2 * δ`. -/ theorem abs_edgeDensity_sub_edgeDensity_le_two_mul (hs : s₂ ⊆ s₁) (ht : t₂ ⊆ t₁) (hδ : 0 ≤ δ) - (hscard : (1 - δ) * s₁.card ≤ s₂.card) (htcard : (1 - δ) * t₁.card ≤ t₂.card) : + (hscard : (1 - δ) * #s₁ ≤ #s₂) (htcard : (1 - δ) * #t₁ ≤ #t₂) : |(edgeDensity r s₂ t₂ : 𝕜) - edgeDensity r s₁ t₁| ≤ 2 * δ := by cases' lt_or_le δ 1 with h h · exact (abs_edgeDensity_sub_edgeDensity_le_two_mul_sub_sq r hs ht hδ h hscard htcard).trans @@ -250,7 +248,7 @@ theorem mk_mem_interedges_comm (hr : Symmetric r) : @swap_mem_interedges_iff _ _ _ _ _ hr (b, a) theorem card_interedges_comm (hr : Symmetric r) (s t : Finset α) : - (interedges r s t).card = (interedges r t s).card := + #(interedges r s t) = #(interedges r t s) := Finset.card_bij (fun (x : α × α) _ ↦ x.swap) (fun _ ↦ (swap_mem_interedges_iff hr).2) (fun _ _ _ _ h ↦ Prod.swap_injective h) fun x h ↦ ⟨x.swap, (swap_mem_interedges_iff hr).2 h, x.swap_swap⟩ @@ -280,16 +278,12 @@ def interedges (s t : Finset α) : Finset (α × α) := def edgeDensity : Finset α → Finset α → ℚ := Rel.edgeDensity G.Adj -theorem interedges_def (s t : Finset α) : - G.interedges s t = (s ×ˢ t).filter fun e ↦ G.Adj e.1 e.2 := - rfl +lemma interedges_def (s t : Finset α) : G.interedges s t = {e ∈ s ×ˢ t | G.Adj e.1 e.2} := rfl -theorem edgeDensity_def (s t : Finset α) : - G.edgeDensity s t = (G.interedges s t).card / (s.card * t.card) := - rfl +lemma edgeDensity_def (s t : Finset α) : G.edgeDensity s t = #(G.interedges s t) / (#s * #t) := rfl theorem card_interedges_div_card (s t : Finset α) : - ((G.interedges s t).card : ℚ) / (s.card * t.card) = G.edgeDensity s t := + (#(G.interedges s t) : ℚ) / (#s * #t) = G.edgeDensity s t := rfl theorem mem_interedges_iff {x : α × α} : x ∈ G.interedges s t ↔ x.1 ∈ s ∧ x.2 ∈ t ∧ G.Adj x.1 x.2 := @@ -331,9 +325,9 @@ theorem interedges_biUnion (s : Finset ι) (t : Finset κ) (f : ι → Finset α Rel.interedges_biUnion _ _ _ _ _ theorem card_interedges_add_card_interedges_compl (h : Disjoint s t) : - (G.interedges s t).card + (Gᶜ.interedges s t).card = s.card * t.card := by + #(G.interedges s t) + #(Gᶜ.interedges s t) = #s * #t := by rw [← card_product, interedges_def, interedges_def] - have : ((s ×ˢ t).filter fun e ↦ Gᶜ.Adj e.1 e.2) = (s ×ˢ t).filter fun e ↦ ¬G.Adj e.1 e.2 := by + have : {e ∈ s ×ˢ t | Gᶜ.Adj e.1 e.2} = {e ∈ s ×ˢ t | ¬G.Adj e.1 e.2} := by refine filter_congr fun x hx ↦ ?_ rw [mem_product] at hx rw [compl_adj, and_iff_right (h.forall_ne_finset hx.1 hx.2)] @@ -349,7 +343,7 @@ theorem edgeDensity_add_edgeDensity_compl (hs : s.Nonempty) (ht : t.Nonempty) (h end DecidableEq -theorem card_interedges_le_mul (s t : Finset α) : (G.interedges s t).card ≤ s.card * t.card := +theorem card_interedges_le_mul (s t : Finset α) : #(G.interedges s t) ≤ #s * #t := Rel.card_interedges_le_mul _ _ _ theorem edgeDensity_nonneg (s t : Finset α) : 0 ≤ G.edgeDensity s t := diff --git a/Mathlib/Combinatorics/SimpleGraph/Finite.lean b/Mathlib/Combinatorics/SimpleGraph/Finite.lean index 3ed8afcfac7a8..e0611556ee07e 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Finite.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Finite.lean @@ -101,27 +101,26 @@ lemma edgeFinset_eq_empty : G.edgeFinset = ∅ ↔ G = ⊥ := by lemma edgeFinset_nonempty : G.edgeFinset.Nonempty ↔ G ≠ ⊥ := by rw [Finset.nonempty_iff_ne_empty, edgeFinset_eq_empty.ne] -theorem edgeFinset_card : G.edgeFinset.card = Fintype.card G.edgeSet := +theorem edgeFinset_card : #G.edgeFinset = Fintype.card G.edgeSet := Set.toFinset_card _ @[simp] -theorem edgeSet_univ_card : (univ : Finset G.edgeSet).card = G.edgeFinset.card := +theorem edgeSet_univ_card : #(univ : Finset G.edgeSet) = #G.edgeFinset := Fintype.card_of_subtype G.edgeFinset fun _ => mem_edgeFinset variable [Fintype V] @[simp] theorem edgeFinset_top [DecidableEq V] : - (⊤ : SimpleGraph V).edgeFinset = univ.filter fun e => ¬e.IsDiag := by - rw [← coe_inj]; simp + (⊤ : SimpleGraph V).edgeFinset = ({e | ¬e.IsDiag} : Finset _) := by simp [← coe_inj] /-- The complete graph on `n` vertices has `n.choose 2` edges. -/ theorem card_edgeFinset_top_eq_card_choose_two [DecidableEq V] : - (⊤ : SimpleGraph V).edgeFinset.card = (Fintype.card V).choose 2 := by + #(⊤ : SimpleGraph V).edgeFinset = (Fintype.card V).choose 2 := by simp_rw [Set.toFinset_card, edgeSet_top, Set.coe_setOf, ← Sym2.card_subtype_not_diag] /-- Any graph on `n` vertices has at most `n.choose 2` edges. -/ -theorem card_edgeFinset_le_card_choose_two : G.edgeFinset.card ≤ (Fintype.card V).choose 2 := by +theorem card_edgeFinset_le_card_choose_two : #G.edgeFinset ≤ (Fintype.card V).choose 2 := by classical rw [← card_edgeFinset_top_eq_card_choose_two] exact card_le_card (edgeFinset_mono le_top) @@ -143,13 +142,13 @@ variable {𝕜 : Type*} [OrderedRing 𝕜] /-- A graph is `r`-*delete-far* from a property `p` if we must delete at least `r` edges from it to get a graph with the property `p`. -/ def DeleteFar (p : SimpleGraph V → Prop) (r : 𝕜) : Prop := - ∀ ⦃s⦄, s ⊆ G.edgeFinset → p (G.deleteEdges s) → r ≤ s.card + ∀ ⦃s⦄, s ⊆ G.edgeFinset → p (G.deleteEdges s) → r ≤ #s variable {G} theorem deleteFar_iff [Fintype (Sym2 V)] : G.DeleteFar p r ↔ ∀ ⦃H : SimpleGraph _⦄ [DecidableRel H.Adj], - H ≤ G → p H → r ≤ G.edgeFinset.card - H.edgeFinset.card := by + H ≤ G → p H → r ≤ #G.edgeFinset - #H.edgeFinset := by classical refine ⟨fun h H _ hHG hH ↦ ?_, fun h s hs hG ↦ ?_⟩ · have := h (sdiff_subset (t := H.edgeFinset)) @@ -203,15 +202,14 @@ theorem singleton_disjoint_neighborFinset : Disjoint {v} (G.neighborFinset v) := Finset.disjoint_singleton_left.mpr <| not_mem_neighborFinset_self _ _ /-- `G.degree v` is the number of vertices adjacent to `v`. -/ -def degree : ℕ := - (G.neighborFinset v).card +def degree : ℕ := #(G.neighborFinset v) -- Porting note: in Lean 3 we could do `simp [← degree]`, but that gives -- "invalid '←' modifier, 'SimpleGraph.degree' is a declaration name to be unfolded". -- In any case, having this lemma is good since there's no guarantee we won't still change -- the definition of `degree`. @[simp] -theorem card_neighborFinset_eq_degree : (G.neighborFinset v).card = G.degree v := rfl +theorem card_neighborFinset_eq_degree : #(G.neighborFinset v) = G.degree v := rfl @[simp] theorem card_neighborSet_eq_degree : Fintype.card (G.neighborSet v) = G.degree v := @@ -240,8 +238,7 @@ theorem card_incidenceSet_eq_degree [DecidableEq V] : simp @[simp] -theorem card_incidenceFinset_eq_degree [DecidableEq V] : - (G.incidenceFinset v).card = G.degree v := by +theorem card_incidenceFinset_eq_degree [DecidableEq V] : #(G.incidenceFinset v) = G.degree v := by rw [← G.card_incidenceSet_eq_degree] apply Set.toFinset_card @@ -251,7 +248,7 @@ theorem mem_incidenceFinset [DecidableEq V] (e : Sym2 V) : Set.mem_toFinset theorem incidenceFinset_eq_filter [DecidableEq V] [Fintype G.edgeSet] : - G.incidenceFinset v = G.edgeFinset.filter (v ∈ ·) := by + G.incidenceFinset v = {e ∈ G.edgeFinset | v ∈ e} := by ext e induction e simp [mk'_mem_incidenceSet_iff] @@ -294,9 +291,7 @@ instance neighborSetFintype [DecidableRel G.Adj] (v : V) : Fintype (G.neighborSe _ theorem neighborFinset_eq_filter {v : V} [DecidableRel G.Adj] : - G.neighborFinset v = Finset.univ.filter (G.Adj v) := by - ext - simp + G.neighborFinset v = ({w | G.Adj v w} : Finset _) := by ext; simp theorem neighborFinset_compl [DecidableEq V] [DecidableRel G.Adj] (v : V) : Gᶜ.neighborFinset v = (G.neighborFinset v)ᶜ \ {v} := by diff --git a/Mathlib/Combinatorics/SimpleGraph/Operations.lean b/Mathlib/Combinatorics/SimpleGraph/Operations.lean index 814a10a5cc143..fea650ca66f09 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Operations.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Operations.lean @@ -34,7 +34,7 @@ variable {G} {W : Type*} {G' : SimpleGraph W} (f : G ≃g G') include f in theorem card_edgeFinset_eq [Fintype G.edgeSet] [Fintype G'.edgeSet] : - G.edgeFinset.card = G'.edgeFinset.card := by + #G.edgeFinset = #G'.edgeFinset := by apply Finset.card_eq_of_equiv simp only [Set.mem_toFinset] exact f.mapEdgeSet @@ -116,7 +116,7 @@ lemma disjoint_sdiff_neighborFinset_image : aesop theorem card_edgeFinset_replaceVertex_of_not_adj (hn : ¬G.Adj s t) : - (G.replaceVertex s t).edgeFinset.card = G.edgeFinset.card + G.degree s - G.degree t := by + #(G.replaceVertex s t).edgeFinset = #G.edgeFinset + G.degree s - G.degree t := by have inc : G.incidenceFinset t ⊆ G.edgeFinset := by simp [incidenceFinset, incidenceSet_subset] rw [G.edgeFinset_replaceVertex_of_not_adj hn, card_union_of_disjoint G.disjoint_sdiff_neighborFinset_image, card_sdiff inc, @@ -127,7 +127,7 @@ theorem card_edgeFinset_replaceVertex_of_not_adj (hn : ¬G.Adj s t) : aesop theorem card_edgeFinset_replaceVertex_of_adj (ha : G.Adj s t) : - (G.replaceVertex s t).edgeFinset.card = G.edgeFinset.card + G.degree s - G.degree t - 1 := by + #(G.replaceVertex s t).edgeFinset = #G.edgeFinset + G.degree s - G.degree t - 1 := by have inc : G.incidenceFinset t ⊆ G.edgeFinset := by simp [incidenceFinset, incidenceSet_subset] rw [G.edgeFinset_replaceVertex_of_adj ha, card_sdiff (by simp [ha]), card_union_of_disjoint G.disjoint_sdiff_neighborFinset_image, card_sdiff inc, @@ -180,7 +180,7 @@ theorem edgeFinset_sup_edge [Fintype (edgeSet (G ⊔ edge s t))] (hn : ¬G.Adj s simp_rw [edgeFinset, edge_edgeSet_of_ne h]; rfl theorem card_edgeFinset_sup_edge [Fintype (edgeSet (G ⊔ edge s t))] (hn : ¬G.Adj s t) (h : s ≠ t) : - (G ⊔ edge s t).edgeFinset.card = G.edgeFinset.card + 1 := by + #(G ⊔ edge s t).edgeFinset = #G.edgeFinset + 1 := by rw [G.edgeFinset_sup_edge hn h, card_cons] end AddEdge diff --git a/Mathlib/Combinatorics/SimpleGraph/Regularity/Bound.lean b/Mathlib/Combinatorics/SimpleGraph/Regularity/Bound.lean index 05fda50393f3e..e70175d8c7398 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Regularity/Bound.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Regularity/Bound.lean @@ -58,9 +58,9 @@ open SzemerediRegularity variable {α : Type*} [DecidableEq α] [Fintype α] {P : Finpartition (univ : Finset α)} {u : Finset α} {ε : ℝ} -local notation3 "m" => (card α / stepBound P.parts.card : ℕ) +local notation3 "m" => (card α / stepBound #P.parts : ℕ) -local notation3 "a" => (card α / P.parts.card - m * 4 ^ P.parts.card : ℕ) +local notation3 "a" => (card α / #P.parts - m * 4 ^ #P.parts : ℕ) namespace SzemerediRegularity.Positivity @@ -68,7 +68,7 @@ private theorem eps_pos {ε : ℝ} {n : ℕ} (h : 100 ≤ (4 : ℝ) ^ n * ε ^ 5 (Odd.pow_pos_iff (by decide)).mp (pos_of_mul_pos_right ((show 0 < (100 : ℝ) by norm_num).trans_le h) (by positivity)) -private theorem m_pos [Nonempty α] (hPα : P.parts.card * 16 ^ P.parts.card ≤ card α) : 0 < m := +private theorem m_pos [Nonempty α] (hPα : #P.parts * 16 ^ #P.parts ≤ card α) : 0 < m := Nat.div_pos ((Nat.mul_le_mul_left _ <| Nat.pow_le_pow_left (by norm_num) _).trans hPα) <| stepBound_pos (P.parts_nonempty <| univ_nonempty.ne_empty).card_pos @@ -101,49 +101,48 @@ namespace SzemerediRegularity open scoped SzemerediRegularity.Positivity -theorem m_pos [Nonempty α] (hPα : P.parts.card * 16 ^ P.parts.card ≤ card α) : 0 < m := by +theorem m_pos [Nonempty α] (hPα : #P.parts * 16 ^ #P.parts ≤ card α) : 0 < m := by sz_positivity theorem coe_m_add_one_pos : 0 < (m : ℝ) + 1 := by positivity -theorem one_le_m_coe [Nonempty α] (hPα : P.parts.card * 16 ^ P.parts.card ≤ card α) : (1 : ℝ) ≤ m := +theorem one_le_m_coe [Nonempty α] (hPα : #P.parts * 16 ^ #P.parts ≤ card α) : (1 : ℝ) ≤ m := Nat.one_le_cast.2 <| m_pos hPα -theorem eps_pow_five_pos (hPε : 100 ≤ (4 : ℝ) ^ P.parts.card * ε ^ 5) : ↑0 < ε ^ 5 := +theorem eps_pow_five_pos (hPε : 100 ≤ (4 : ℝ) ^ #P.parts * ε ^ 5) : ↑0 < ε ^ 5 := pos_of_mul_pos_right ((by norm_num : (0 : ℝ) < 100).trans_le hPε) <| pow_nonneg (by norm_num) _ -theorem eps_pos (hPε : 100 ≤ (4 : ℝ) ^ P.parts.card * ε ^ 5) : 0 < ε := +theorem eps_pos (hPε : 100 ≤ (4 : ℝ) ^ #P.parts * ε ^ 5) : 0 < ε := (Odd.pow_pos_iff (by decide)).mp (eps_pow_five_pos hPε) -theorem hundred_div_ε_pow_five_le_m [Nonempty α] (hPα : P.parts.card * 16 ^ P.parts.card ≤ card α) - (hPε : 100 ≤ (4 : ℝ) ^ P.parts.card * ε ^ 5) : 100 / ε ^ 5 ≤ m := +theorem hundred_div_ε_pow_five_le_m [Nonempty α] (hPα : #P.parts * 16 ^ #P.parts ≤ card α) + (hPε : 100 ≤ (4 : ℝ) ^ #P.parts * ε ^ 5) : 100 / ε ^ 5 ≤ m := (div_le_of_le_mul₀ (eps_pow_five_pos hPε).le (by positivity) hPε).trans <| by norm_cast rwa [Nat.le_div_iff_mul_le' (stepBound_pos (P.parts_nonempty <| univ_nonempty.ne_empty).card_pos), stepBound, mul_left_comm, ← mul_pow] -theorem hundred_le_m [Nonempty α] (hPα : P.parts.card * 16 ^ P.parts.card ≤ card α) - (hPε : 100 ≤ (4 : ℝ) ^ P.parts.card * ε ^ 5) (hε : ε ≤ 1) : 100 ≤ m := +theorem hundred_le_m [Nonempty α] (hPα : #P.parts * 16 ^ #P.parts ≤ card α) + (hPε : 100 ≤ (4 : ℝ) ^ #P.parts * ε ^ 5) (hε : ε ≤ 1) : 100 ≤ m := mod_cast (hundred_div_ε_pow_five_le_m hPα hPε).trans' (le_div_self (by norm_num) (by sz_positivity) <| pow_le_one₀ (by sz_positivity) hε) -theorem a_add_one_le_four_pow_parts_card : a + 1 ≤ 4 ^ P.parts.card := by - have h : 1 ≤ 4 ^ P.parts.card := one_le_pow₀ (by norm_num) +theorem a_add_one_le_four_pow_parts_card : a + 1 ≤ 4 ^ #P.parts := by + have h : 1 ≤ 4 ^ #P.parts := one_le_pow₀ (by norm_num) rw [stepBound, ← Nat.div_div_eq_div_mul] conv_rhs => rw [← Nat.sub_add_cancel h] rw [add_le_add_iff_right, tsub_le_iff_left, ← Nat.add_sub_assoc h] exact Nat.le_sub_one_of_lt (Nat.lt_div_mul_add h) -theorem card_aux₁ (hucard : u.card = m * 4 ^ P.parts.card + a) : - (4 ^ P.parts.card - a) * m + a * (m + 1) = u.card := by +theorem card_aux₁ (hucard : #u = m * 4 ^ #P.parts + a) : + (4 ^ #P.parts - a) * m + a * (m + 1) = #u := by rw [hucard, mul_add, mul_one, ← add_assoc, ← add_mul, Nat.sub_add_cancel ((Nat.le_succ _).trans a_add_one_le_four_pow_parts_card), mul_comm] -theorem card_aux₂ (hP : P.IsEquipartition) (hu : u ∈ P.parts) - (hucard : ¬u.card = m * 4 ^ P.parts.card + a) : - (4 ^ P.parts.card - (a + 1)) * m + (a + 1) * (m + 1) = u.card := by - have : m * 4 ^ P.parts.card ≤ card α / P.parts.card := by +theorem card_aux₂ (hP : P.IsEquipartition) (hu : u ∈ P.parts) (hucard : #u ≠ m * 4 ^ #P.parts + a) : + (4 ^ #P.parts - (a + 1)) * m + (a + 1) * (m + 1) = #u := by + have : m * 4 ^ #P.parts ≤ card α / #P.parts := by rw [stepBound, ← Nat.div_div_eq_div_mul] exact Nat.div_mul_le_self _ _ rw [Nat.add_sub_of_le this] at hucard @@ -152,7 +151,7 @@ theorem card_aux₂ (hP : P.IsEquipartition) (hu : u ∈ P.parts) Nat.add_sub_of_le this, card_univ] theorem pow_mul_m_le_card_part (hP : P.IsEquipartition) (hu : u ∈ P.parts) : - (4 : ℝ) ^ P.parts.card * m ≤ u.card := by + (4 : ℝ) ^ #P.parts * m ≤ #u := by norm_cast rw [stepBound, ← Nat.div_div_eq_div_mul] exact (Nat.mul_div_le _ _).trans (hP.average_le_card_part hu) @@ -200,35 +199,34 @@ theorem bound_pos : 0 < bound ε l := variable {ι 𝕜 : Type*} [LinearOrderedField 𝕜] (r : ι → ι → Prop) [DecidableRel r] {s t : Finset ι} {x : 𝕜} -theorem mul_sq_le_sum_sq (hst : s ⊆ t) (f : ι → 𝕜) (hs : x ^ 2 ≤ ((∑ i ∈ s, f i) / s.card) ^ 2) - (hs' : (s.card : 𝕜) ≠ 0) : (s.card : 𝕜) * x ^ 2 ≤ ∑ i ∈ t, f i ^ 2 := +theorem mul_sq_le_sum_sq (hst : s ⊆ t) (f : ι → 𝕜) (hs : x ^ 2 ≤ ((∑ i ∈ s, f i) / #s) ^ 2) + (hs' : (#s : 𝕜) ≠ 0) : (#s : 𝕜) * x ^ 2 ≤ ∑ i ∈ t, f i ^ 2 := (mul_le_mul_of_nonneg_left (hs.trans sum_div_card_sq_le_sum_sq_div_card) <| Nat.cast_nonneg _).trans <| (mul_div_cancel₀ _ hs').le.trans <| sum_le_sum_of_subset_of_nonneg hst fun _ _ _ => sq_nonneg _ theorem add_div_le_sum_sq_div_card (hst : s ⊆ t) (f : ι → 𝕜) (d : 𝕜) (hx : 0 ≤ x) - (hs : x ≤ |(∑ i ∈ s, f i) / s.card - (∑ i ∈ t, f i) / t.card|) - (ht : d ≤ ((∑ i ∈ t, f i) / t.card) ^ 2) : - d + s.card / t.card * x ^ 2 ≤ (∑ i ∈ t, f i ^ 2) / t.card := by - obtain hscard | hscard := (s.card.cast_nonneg : (0 : 𝕜) ≤ s.card).eq_or_lt + (hs : x ≤ |(∑ i ∈ s, f i) / #s - (∑ i ∈ t, f i) / #t|) (ht : d ≤ ((∑ i ∈ t, f i) / #t) ^ 2) : + d + #s / #t * x ^ 2 ≤ (∑ i ∈ t, f i ^ 2) / #t := by + obtain hscard | hscard := ((#s).cast_nonneg : (0 : 𝕜) ≤ #s).eq_or_lt · simpa [← hscard] using ht.trans sum_div_card_sq_le_sum_sq_div_card - have htcard : (0 : 𝕜) < t.card := hscard.trans_le (Nat.cast_le.2 (card_le_card hst)) - have h₁ : x ^ 2 ≤ ((∑ i ∈ s, f i) / s.card - (∑ i ∈ t, f i) / t.card) ^ 2 := + have htcard : (0 : 𝕜) < #t := hscard.trans_le (Nat.cast_le.2 (card_le_card hst)) + have h₁ : x ^ 2 ≤ ((∑ i ∈ s, f i) / #s - (∑ i ∈ t, f i) / #t) ^ 2 := sq_le_sq.2 (by rwa [abs_of_nonneg hx]) - have h₂ : x ^ 2 ≤ ((∑ i ∈ s, (f i - (∑ j ∈ t, f j) / t.card)) / s.card) ^ 2 := by + have h₂ : x ^ 2 ≤ ((∑ i ∈ s, (f i - (∑ j ∈ t, f j) / #t)) / #s) ^ 2 := by apply h₁.trans rw [sum_sub_distrib, sum_const, nsmul_eq_mul, sub_div, mul_div_cancel_left₀ _ hscard.ne'] apply (add_le_add_right ht _).trans rw [← mul_div_right_comm, le_div_iff₀ htcard, add_mul, div_mul_cancel₀ _ htcard.ne'] - have h₃ := mul_sq_le_sum_sq hst (fun i => (f i - (∑ j ∈ t, f j) / t.card)) h₂ hscard.ne' + have h₃ := mul_sq_le_sum_sq hst (fun i => (f i - (∑ j ∈ t, f j) / #t)) h₂ hscard.ne' apply (add_le_add_left h₃ _).trans -- Porting note: was - -- `simp [← mul_div_right_comm _ (t.card : 𝕜), sub_div' _ _ _ htcard.ne', ← sum_div, ← add_div,` + -- `simp [← mul_div_right_comm _ (#t : 𝕜), sub_div' _ _ _ htcard.ne', ← sum_div, ← add_div,` -- ` mul_pow, div_le_iff₀ (sq_pos_of_ne_zero htcard.ne'), sub_sq, sum_add_distrib, ← sum_mul,` -- ` ← mul_sum]` simp_rw [sub_div' _ _ _ htcard.ne'] conv_lhs => enter [2, 2, x]; rw [div_pow] - rw [div_pow, ← sum_div, ← mul_div_right_comm _ (t.card : 𝕜), ← add_div, + rw [div_pow, ← sum_div, ← mul_div_right_comm _ (#t : 𝕜), ← add_div, div_le_iff₀ (sq_pos_of_ne_zero htcard.ne')] simp_rw [sub_sq, sum_add_distrib, sum_const, nsmul_eq_mul, sum_sub_distrib, mul_pow, ← sum_mul, ← mul_sum, ← sum_mul] diff --git a/Mathlib/Combinatorics/SimpleGraph/Regularity/Chunk.lean b/Mathlib/Combinatorics/SimpleGraph/Regularity/Chunk.lean index 14c2df9eb7249..07a903fca9f50 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Regularity/Chunk.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Regularity/Chunk.lean @@ -44,7 +44,7 @@ variable {α : Type*} [Fintype α] [DecidableEq α] {P : Finpartition (univ : Fi (hP : P.IsEquipartition) (G : SimpleGraph α) [DecidableRel G.Adj] (ε : ℝ) {U : Finset α} (hU : U ∈ P.parts) (V : Finset α) -local notation3 "m" => (card α / stepBound P.parts.card : ℕ) +local notation3 "m" => (card α / stepBound #P.parts : ℕ) /-! ### Definitions @@ -56,16 +56,16 @@ contained in the corresponding witness of non-uniformity. /-- The portion of `SzemerediRegularity.increment` which partitions `U`. -/ noncomputable def chunk : Finpartition U := - if hUcard : U.card = m * 4 ^ P.parts.card + (card α / P.parts.card - m * 4 ^ P.parts.card) then + if hUcard : #U = m * 4 ^ #P.parts + (card α / #P.parts - m * 4 ^ #P.parts) then (atomise U <| P.nonuniformWitnesses G ε U).equitabilise <| card_aux₁ hUcard else (atomise U <| P.nonuniformWitnesses G ε U).equitabilise <| card_aux₂ hP hU hUcard -- `hP` and `hU` are used to get that `U` has size --- `m * 4 ^ P.parts.card + a or m * 4 ^ P.parts.card + a + 1` +-- `m * 4 ^ #P.parts + a or m * 4 ^ #P.parts + a + 1` /-- The portion of `SzemerediRegularity.chunk` which is contained in the witness of non-uniformity of `U` and `V`. -/ noncomputable def star (V : Finset α) : Finset (Finset α) := - (chunk hP G ε hU).parts.filter (· ⊆ G.nonuniformWitness ε U V) + {A ∈ (chunk hP G ε hU).parts | A ⊆ G.nonuniformWitness ε U V} /-! ### Density estimates @@ -85,14 +85,13 @@ theorem star_subset_chunk : star hP G ε hU V ⊆ (chunk hP G ε hU).parts := private theorem card_nonuniformWitness_sdiff_biUnion_star (hV : V ∈ P.parts) (hUV : U ≠ V) (h₂ : ¬G.IsUniform ε U V) : - (G.nonuniformWitness ε U V \ (star hP G ε hU V).biUnion id).card ≤ - 2 ^ (P.parts.card - 1) * m := by + #(G.nonuniformWitness ε U V \ (star hP G ε hU V).biUnion id) ≤ 2 ^ (#P.parts - 1) * m := by have hX : G.nonuniformWitness ε U V ∈ P.nonuniformWitnesses G ε U := nonuniformWitness_mem_nonuniformWitnesses h₂ hV hUV have q : G.nonuniformWitness ε U V \ (star hP G ε hU V).biUnion id ⊆ - ((atomise U <| P.nonuniformWitnesses G ε U).parts.filter fun B => - B ⊆ G.nonuniformWitness ε U V ∧ B.Nonempty).biUnion - fun B => B \ ((chunk hP G ε hU).parts.filter (· ⊆ B)).biUnion id := by + {B ∈ (atomise U <| P.nonuniformWitnesses G ε U).parts | + B ⊆ G.nonuniformWitness ε U V ∧ B.Nonempty}.biUnion + fun B => B \ {A ∈ (chunk hP G ε hU).parts | A ⊆ B}.biUnion id := by intro x hx rw [← biUnion_filter_atomise hX (G.nonuniformWitness_subset h₂), star, mem_sdiff, mem_biUnion] at hx @@ -101,10 +100,10 @@ private theorem card_nonuniformWitness_sdiff_biUnion_star (hV : V ∈ P.parts) ( obtain ⟨⟨B, hB₁, hB₂⟩, hx⟩ := hx exact ⟨B, hB₁, hB₂, fun A hA AB => hx A hA <| AB.trans hB₁.2.1⟩ apply (card_le_card q).trans (card_biUnion_le.trans _) - trans ∑ _i in (atomise U <| P.nonuniformWitnesses G ε U).parts.filter fun B => + trans ∑ B ∈ (atomise U <| P.nonuniformWitnesses G ε U).parts with B ⊆ G.nonuniformWitness ε U V ∧ B.Nonempty, m · suffices ∀ B ∈ (atomise U <| P.nonuniformWitnesses G ε U).parts, - (B \ ((chunk hP G ε hU).parts.filter (· ⊆ B)).biUnion id).card ≤ m by + #(B \ {A ∈ (chunk hP G ε hU).parts | A ⊆ B}.biUnion id) ≤ m by exact sum_le_sum fun B hB => this B <| filter_subset _ _ hB intro B hB unfold chunk @@ -118,43 +117,43 @@ private theorem card_nonuniformWitness_sdiff_biUnion_star (hV : V ∈ P.parts) ( 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) - (hUV : U ≠ V) (hunif : ¬G.IsUniform ε U V) (hPε : ↑100 ≤ ↑4 ^ P.parts.card * ε ^ 5) + (hUV : U ≠ V) (hunif : ¬G.IsUniform ε U V) (hPε : ↑100 ≤ ↑4 ^ #P.parts * ε ^ 5) (hε₁ : ε ≤ 1) : - (1 - ε / 10) * (G.nonuniformWitness ε U V).card ≤ ((star hP G ε hU V).biUnion id).card := by - have hP₁ : 0 < P.parts.card := Finset.card_pos.2 ⟨_, hU⟩ - have : (↑2 ^ P.parts.card : ℝ) * m / (U.card * ε) ≤ ε / 10 := by + (1 - ε / 10) * #(G.nonuniformWitness ε U V) ≤ #((star hP G ε hU V).biUnion id) := by + have hP₁ : 0 < #P.parts := Finset.card_pos.2 ⟨_, hU⟩ + have : (↑2 ^ #P.parts : ℝ) * m / (#U * ε) ≤ ε / 10 := by rw [← div_div, div_le_iff₀'] swap · sz_positivity - refine le_of_mul_le_mul_left ?_ (pow_pos zero_lt_two P.parts.card) + refine le_of_mul_le_mul_left ?_ (pow_pos zero_lt_two #P.parts) calc - ↑2 ^ P.parts.card * ((↑2 ^ P.parts.card * m : ℝ) / U.card) = - ((2 : ℝ) * 2) ^ P.parts.card * m / U.card := by + ↑2 ^ #P.parts * ((↑2 ^ #P.parts * m : ℝ) / #U) = + ((2 : ℝ) * 2) ^ #P.parts * m / #U := by rw [mul_pow, ← mul_div_assoc, mul_assoc] - _ = ↑4 ^ P.parts.card * m / U.card := by norm_num + _ = ↑4 ^ #P.parts * m / #U := by norm_num _ ≤ 1 := div_le_one_of_le₀ (pow_mul_m_le_card_part hP hU) (cast_nonneg _) - _ ≤ ↑2 ^ P.parts.card * ε ^ 2 / 10 := by + _ ≤ ↑2 ^ #P.parts * ε ^ 2 / 10 := by refine (one_le_sq_iff <| by positivity).1 ?_ rw [div_pow, mul_pow, pow_right_comm, ← pow_mul ε, one_le_div (sq_pos_of_ne_zero <| by norm_num)] calc (↑10 ^ 2) = 100 := by norm_num - _ ≤ ↑4 ^ P.parts.card * ε ^ 5 := hPε - _ ≤ ↑4 ^ P.parts.card * ε ^ 4 := + _ ≤ ↑4 ^ #P.parts * ε ^ 5 := hPε + _ ≤ ↑4 ^ #P.parts * ε ^ 4 := (mul_le_mul_of_nonneg_left (pow_le_pow_of_le_one (by sz_positivity) hε₁ <| le_succ _) (by positivity)) - _ = (↑2 ^ 2) ^ P.parts.card * ε ^ (2 * 2) := by norm_num - _ = ↑2 ^ P.parts.card * (ε * (ε / 10)) := by rw [mul_div_assoc, sq, mul_div_assoc] + _ = (↑2 ^ 2) ^ #P.parts * ε ^ (2 * 2) := by norm_num + _ = ↑2 ^ #P.parts * (ε * (ε / 10)) := by rw [mul_div_assoc, sq, mul_div_assoc] calc - (↑1 - ε / 10) * (G.nonuniformWitness ε U V).card ≤ - (↑1 - ↑2 ^ P.parts.card * m / (U.card * ε)) * (G.nonuniformWitness ε U V).card := + (↑1 - ε / 10) * #(G.nonuniformWitness ε U V) ≤ + (↑1 - ↑2 ^ #P.parts * m / (#U * ε)) * #(G.nonuniformWitness ε U V) := mul_le_mul_of_nonneg_right (sub_le_sub_left this _) (cast_nonneg _) - _ = (G.nonuniformWitness ε U V).card - - ↑2 ^ P.parts.card * m / (U.card * ε) * (G.nonuniformWitness ε U V).card := by + _ = #(G.nonuniformWitness ε U V) - + ↑2 ^ #P.parts * m / (#U * ε) * #(G.nonuniformWitness ε U V) := by rw [sub_mul, one_mul] - _ ≤ (G.nonuniformWitness ε U V).card - ↑2 ^ (P.parts.card - 1) * m := by + _ ≤ #(G.nonuniformWitness ε U V) - ↑2 ^ (#P.parts - 1) * m := by refine sub_le_sub_left ?_ _ - have : (2 : ℝ) ^ P.parts.card = ↑2 ^ (P.parts.card - 1) * 2 := by + have : (2 : ℝ) ^ #P.parts = ↑2 ^ (#P.parts - 1) * 2 := by rw [← _root_.pow_succ, tsub_add_cancel_of_le (succ_le_iff.2 hP₁)] rw [← mul_div_right_comm, this, mul_right_comm _ (2 : ℝ), mul_assoc, le_div_iff₀] · refine mul_le_mul_of_nonneg_left ?_ (by positivity) @@ -162,7 +161,7 @@ private theorem one_sub_eps_mul_card_nonuniformWitness_le_card_star (hV : V ∈ (le_mul_of_one_le_left (cast_nonneg _) one_le_two) have := Finset.card_pos.mpr (P.nonempty_of_mem_parts hU) sz_positivity - _ ≤ ((star hP G ε hU V).biUnion id).card := by + _ ≤ #((star hP G ε hU V).biUnion id) := by rw [sub_le_comm, ← cast_sub (card_le_card <| biUnion_star_subset_nonuniformWitness hP G ε hU V), ← card_sdiff (biUnion_star_subset_nonuniformWitness hP G ε hU V)] @@ -171,7 +170,7 @@ private theorem one_sub_eps_mul_card_nonuniformWitness_le_card_star (hV : V ∈ /-! ### `chunk` -/ -theorem card_chunk (hm : m ≠ 0) : (chunk hP G ε hU).parts.card = 4 ^ P.parts.card := by +theorem card_chunk (hm : m ≠ 0) : #(chunk hP G ε hU).parts = 4 ^ #P.parts := by unfold chunk split_ifs · rw [card_parts_equitabilise _ _ hm, tsub_add_cancel_of_le] @@ -179,23 +178,23 @@ theorem card_chunk (hm : m ≠ 0) : (chunk hP G ε hU).parts.card = 4 ^ P.parts. · rw [card_parts_equitabilise _ _ hm, tsub_add_cancel_of_le a_add_one_le_four_pow_parts_card] theorem card_eq_of_mem_parts_chunk (hs : s ∈ (chunk hP G ε hU).parts) : - s.card = m ∨ s.card = m + 1 := by + #s = m ∨ #s = m + 1 := by unfold chunk at hs split_ifs at hs <;> exact card_eq_of_mem_parts_equitabilise hs -theorem m_le_card_of_mem_chunk_parts (hs : s ∈ (chunk hP G ε hU).parts) : m ≤ s.card := +theorem m_le_card_of_mem_chunk_parts (hs : s ∈ (chunk hP G ε hU).parts) : m ≤ #s := (card_eq_of_mem_parts_chunk hs).elim ge_of_eq fun i => by simp [i] -theorem card_le_m_add_one_of_mem_chunk_parts (hs : s ∈ (chunk hP G ε hU).parts) : s.card ≤ m + 1 := +theorem card_le_m_add_one_of_mem_chunk_parts (hs : s ∈ (chunk hP G ε hU).parts) : #s ≤ m + 1 := (card_eq_of_mem_parts_chunk hs).elim (fun i => by simp [i]) fun i => i.le theorem card_biUnion_star_le_m_add_one_card_star_mul : - (((star hP G ε hU V).biUnion id).card : ℝ) ≤ (star hP G ε hU V).card * (m + 1) := + (#((star hP G ε hU V).biUnion id) : ℝ) ≤ #(star hP G ε hU V) * (m + 1) := mod_cast card_biUnion_le_card_mul _ _ _ fun _ hs => card_le_m_add_one_of_mem_chunk_parts <| star_subset_chunk hs private theorem le_sum_card_subset_chunk_parts (h𝒜 : 𝒜 ⊆ (chunk hP G ε hU).parts) (hs : s ∈ 𝒜) : - (𝒜.card : ℝ) * s.card * (m / (m + 1)) ≤ (𝒜.sup id).card := by + (#𝒜 : ℝ) * #s * (m / (m + 1)) ≤ #(𝒜.sup id) := by rw [mul_div_assoc', div_le_iff₀ coe_m_add_one_pos, mul_right_comm] refine mul_le_mul ?_ ?_ (cast_nonneg _) (cast_nonneg _) · rw [← (ofSubset _ h𝒜 rfl).sum_card_parts, ofSubset_parts, ← cast_mul, cast_le] @@ -204,7 +203,7 @@ private theorem le_sum_card_subset_chunk_parts (h𝒜 : 𝒜 ⊆ (chunk hP G ε private theorem sum_card_subset_chunk_parts_le (m_pos : (0 : ℝ) < m) (h𝒜 : 𝒜 ⊆ (chunk hP G ε hU).parts) (hs : s ∈ 𝒜) : - ((𝒜.sup id).card : ℝ) ≤ 𝒜.card * s.card * ((m + 1) / m) := by + (#(𝒜.sup id) : ℝ) ≤ #𝒜 * #s * ((m + 1) / m) := by rw [sup_eq_biUnion, mul_div_assoc', le_div_iff₀ m_pos, mul_right_comm] refine mul_le_mul ?_ ?_ (cast_nonneg _) (by positivity) · norm_cast @@ -213,7 +212,7 @@ private theorem sum_card_subset_chunk_parts_le (m_pos : (0 : ℝ) < m) · exact mod_cast m_le_card_of_mem_chunk_parts (h𝒜 hs) private theorem one_sub_le_m_div_m_add_one_sq [Nonempty α] - (hPα : P.parts.card * 16 ^ P.parts.card ≤ card α) (hPε : ↑100 ≤ ↑4 ^ P.parts.card * ε ^ 5) : + (hPα : #P.parts * 16 ^ #P.parts ≤ card α) (hPε : ↑100 ≤ ↑4 ^ #P.parts * ε ^ 5) : ↑1 - ε ^ 5 / ↑50 ≤ (m / (m + 1 : ℝ)) ^ 2 := by have : (m : ℝ) / (m + 1) = 1 - 1 / (m + 1) := by rw [one_sub_div coe_m_add_one_pos.ne', add_sub_cancel_right] @@ -227,7 +226,7 @@ private theorem one_sub_le_m_div_m_add_one_sq [Nonempty α] sz_positivity private theorem m_add_one_div_m_le_one_add [Nonempty α] - (hPα : P.parts.card * 16 ^ P.parts.card ≤ card α) (hPε : ↑100 ≤ ↑4 ^ P.parts.card * ε ^ 5) + (hPα : #P.parts * 16 ^ #P.parts ≤ card α) (hPε : ↑100 ≤ ↑4 ^ #P.parts * ε ^ 5) (hε₁ : ε ≤ 1) : ((m + 1 : ℝ) / m) ^ 2 ≤ ↑1 + ε ^ 5 / 49 := by rw [same_add_div] swap; · sz_positivity @@ -243,11 +242,11 @@ private theorem m_add_one_div_m_le_one_add [Nonempty α] exact (pow_le_one₀ (by sz_positivity) hε₁).trans (by norm_num) private theorem density_sub_eps_le_sum_density_div_card [Nonempty α] - (hPα : P.parts.card * 16 ^ P.parts.card ≤ card α) (hPε : ↑100 ≤ ↑4 ^ P.parts.card * ε ^ 5) + (hPα : #P.parts * 16 ^ #P.parts ≤ card α) (hPε : ↑100 ≤ ↑4 ^ #P.parts * ε ^ 5) {hU : U ∈ P.parts} {hV : V ∈ P.parts} {A B : Finset (Finset α)} (hA : A ⊆ (chunk hP G ε hU).parts) (hB : B ⊆ (chunk hP G ε hV).parts) : (G.edgeDensity (A.biUnion id) (B.biUnion id)) - ε ^ 5 / 50 ≤ - (∑ ab ∈ A.product B, (G.edgeDensity ab.1 ab.2 : ℝ)) / (A.card * B.card) := by + (∑ ab ∈ A.product B, (G.edgeDensity ab.1 ab.2 : ℝ)) / (#A * #B) := by have : ↑(G.edgeDensity (A.biUnion id) (B.biUnion id)) - ε ^ 5 / ↑50 ≤ (↑1 - ε ^ 5 / 50) * G.edgeDensity (A.biUnion id) (B.biUnion id) := by rw [sub_mul, one_mul, sub_le_sub_iff_left] @@ -266,7 +265,7 @@ private theorem density_sub_eps_le_sum_density_div_card [Nonempty α] apply sum_le_sum simp only [and_imp, Prod.forall, mem_product] rintro x y hx hy - rw [mul_mul_mul_comm, mul_comm (x.card : ℝ), mul_comm (y.card : ℝ), le_div_iff₀, mul_assoc] + rw [mul_mul_mul_comm, mul_comm (#x : ℝ), mul_comm (#y : ℝ), le_div_iff₀, mul_assoc] · refine mul_le_of_le_one_right (cast_nonneg _) ?_ rw [div_mul_eq_mul_div, ← mul_assoc, mul_assoc] refine div_le_one_of_le₀ ?_ (by positivity) @@ -281,10 +280,10 @@ private theorem density_sub_eps_le_sum_density_div_card [Nonempty α] exacts [⟨_, hx⟩, nonempty_of_mem_parts _ (hA hx), ⟨_, hy⟩, nonempty_of_mem_parts _ (hB hy)] private theorem sum_density_div_card_le_density_add_eps [Nonempty α] - (hPα : P.parts.card * 16 ^ P.parts.card ≤ card α) (hPε : ↑100 ≤ ↑4 ^ P.parts.card * ε ^ 5) + (hPα : #P.parts * 16 ^ #P.parts ≤ card α) (hPε : ↑100 ≤ ↑4 ^ #P.parts * ε ^ 5) (hε₁ : ε ≤ 1) {hU : U ∈ P.parts} {hV : V ∈ P.parts} {A B : Finset (Finset α)} (hA : A ⊆ (chunk hP G ε hU).parts) (hB : B ⊆ (chunk hP G ε hV).parts) : - (∑ ab ∈ A.product B, G.edgeDensity ab.1 ab.2 : ℝ) / (A.card * B.card) ≤ + (∑ ab ∈ A.product B, G.edgeDensity ab.1 ab.2 : ℝ) / (#A * #B) ≤ G.edgeDensity (A.biUnion id) (B.biUnion id) + ε ^ 5 / 49 := by have : (↑1 + ε ^ 5 / ↑49) * G.edgeDensity (A.biUnion id) (B.biUnion id) ≤ G.edgeDensity (A.biUnion id) (B.biUnion id) + ε ^ 5 / 49 := by @@ -303,7 +302,7 @@ private theorem sum_density_div_card_le_density_add_eps [Nonempty α] apply sum_le_sum simp only [and_imp, Prod.forall, mem_product, show A.product B = A ×ˢ B by rfl] intro x y hx hy - rw [mul_mul_mul_comm, mul_comm (x.card : ℝ), mul_comm (y.card : ℝ), div_le_iff₀, mul_assoc] + rw [mul_mul_mul_comm, mul_comm (#x : ℝ), mul_comm (#y : ℝ), div_le_iff₀, mul_assoc] · refine le_mul_of_one_le_right (cast_nonneg _) ?_ rw [div_mul_eq_mul_div, one_le_div] · refine le_trans ?_ (mul_le_mul_of_nonneg_right (m_add_one_div_m_le_one_add hPα hPε hε₁) ?_) @@ -319,28 +318,28 @@ private theorem sum_density_div_card_le_density_add_eps [Nonempty α] exacts [⟨_, hx⟩, nonempty_of_mem_parts _ (hA hx), ⟨_, hy⟩, nonempty_of_mem_parts _ (hB hy)] private theorem average_density_near_total_density [Nonempty α] - (hPα : P.parts.card * 16 ^ P.parts.card ≤ card α) (hPε : ↑100 ≤ ↑4 ^ P.parts.card * ε ^ 5) + (hPα : #P.parts * 16 ^ #P.parts ≤ card α) (hPε : ↑100 ≤ ↑4 ^ #P.parts * ε ^ 5) (hε₁ : ε ≤ 1) {hU : U ∈ P.parts} {hV : V ∈ P.parts} {A B : Finset (Finset α)} (hA : A ⊆ (chunk hP G ε hU).parts) (hB : B ⊆ (chunk hP G ε hV).parts) : - |(∑ ab ∈ A.product B, G.edgeDensity ab.1 ab.2 : ℝ) / (A.card * B.card) - + |(∑ ab ∈ A.product B, G.edgeDensity ab.1 ab.2 : ℝ) / (#A * #B) - G.edgeDensity (A.biUnion id) (B.biUnion id)| ≤ ε ^ 5 / 49 := by rw [abs_sub_le_iff] constructor · rw [sub_le_iff_le_add'] exact sum_density_div_card_le_density_add_eps hPα hPε hε₁ hA hB suffices (G.edgeDensity (A.biUnion id) (B.biUnion id) : ℝ) - - (∑ ab ∈ A.product B, (G.edgeDensity ab.1 ab.2 : ℝ)) / (A.card * B.card) ≤ ε ^ 5 / 50 by + (∑ ab ∈ A.product B, (G.edgeDensity ab.1 ab.2 : ℝ)) / (#A * #B) ≤ ε ^ 5 / 50 by apply this.trans gcongr <;> [sz_positivity; norm_num] rw [sub_le_iff_le_add, ← sub_le_iff_le_add'] apply density_sub_eps_le_sum_density_div_card hPα hPε hA hB private theorem edgeDensity_chunk_aux [Nonempty α] - (hPα : P.parts.card * 16 ^ P.parts.card ≤ card α) (hPε : ↑100 ≤ ↑4 ^ P.parts.card * ε ^ 5) + (hPα : #P.parts * 16 ^ #P.parts ≤ card α) (hPε : ↑100 ≤ ↑4 ^ #P.parts * ε ^ 5) (hU : U ∈ P.parts) (hV : V ∈ P.parts) : (G.edgeDensity U V : ℝ) ^ 2 - ε ^ 5 / ↑25 ≤ ((∑ ab ∈ (chunk hP G ε hU).parts.product (chunk hP G ε hV).parts, - (G.edgeDensity ab.1 ab.2 : ℝ)) / ↑16 ^ P.parts.card) ^ 2 := by + (G.edgeDensity ab.1 ab.2 : ℝ)) / ↑16 ^ #P.parts) ^ 2 := by obtain hGε | hGε := le_total (G.edgeDensity U V : ℝ) (ε ^ 5 / 50) · refine (sub_nonpos_of_le <| (sq_le ?_ ?_).trans <| hGε.trans ?_).trans (sq_nonneg _) · exact mod_cast G.edgeDensity_nonneg _ _ @@ -349,7 +348,7 @@ private theorem edgeDensity_chunk_aux [Nonempty α] rw [← sub_nonneg] at hGε have : ↑(G.edgeDensity U V) - ε ^ 5 / ↑50 ≤ (∑ ab ∈ (chunk hP G ε hU).parts.product (chunk hP G ε hV).parts, - (G.edgeDensity ab.1 ab.2 : ℝ)) / ↑16 ^ P.parts.card := by + (G.edgeDensity ab.1 ab.2 : ℝ)) / ↑16 ^ #P.parts := by have rflU := Set.Subset.refl (chunk hP G ε hU).parts.toSet have rflV := Set.Subset.refl (chunk hP G ε hV).parts.toSet refine (le_trans ?_ <| density_sub_eps_le_sum_density_div_card hPα hPε rflU rflV).trans ?_ @@ -363,7 +362,7 @@ private theorem edgeDensity_chunk_aux [Nonempty α] show (2 : ℝ) / 50 = 25⁻¹ by norm_num] exact mul_le_of_le_one_right (by sz_positivity) (mod_cast G.edgeDensity_le_one _ _) -private theorem abs_density_star_sub_density_le_eps (hPε : ↑100 ≤ ↑4 ^ P.parts.card * ε ^ 5) +private theorem abs_density_star_sub_density_le_eps (hPε : ↑100 ≤ ↑4 ^ #P.parts * ε ^ 5) (hε₁ : ε ≤ 1) {hU : U ∈ P.parts} {hV : V ∈ P.parts} (hUV' : U ≠ V) (hUV : ¬G.IsUniform ε U V) : |(G.edgeDensity ((star hP G ε hU V).biUnion id) ((star hP G ε hV U).biUnion id) : ℝ) - G.edgeDensity (G.nonuniformWitness ε U V) (G.nonuniformWitness ε V U)| ≤ ε / 5 := by @@ -375,33 +374,33 @@ private theorem abs_density_star_sub_density_le_eps (hPε : ↑100 ≤ ↑4 ^ P. hPε hε₁) using 1 linarith -private theorem eps_le_card_star_div [Nonempty α] (hPα : P.parts.card * 16 ^ P.parts.card ≤ card α) - (hPε : ↑100 ≤ ↑4 ^ P.parts.card * ε ^ 5) (hε₁ : ε ≤ 1) (hU : U ∈ P.parts) (hV : V ∈ P.parts) +private theorem eps_le_card_star_div [Nonempty α] (hPα : #P.parts * 16 ^ #P.parts ≤ card α) + (hPε : ↑100 ≤ ↑4 ^ #P.parts * ε ^ 5) (hε₁ : ε ≤ 1) (hU : U ∈ P.parts) (hV : V ∈ P.parts) (hUV : U ≠ V) (hunif : ¬G.IsUniform ε U V) : - ↑4 / ↑5 * ε ≤ (star hP G ε hU V).card / ↑4 ^ P.parts.card := by + ↑4 / ↑5 * ε ≤ #(star hP G ε hU V) / ↑4 ^ #P.parts := by have hm : (0 : ℝ) ≤ 1 - (↑m)⁻¹ := sub_nonneg_of_le (inv_le_one_of_one_le₀ <| one_le_m_coe hPα) have hε : 0 ≤ 1 - ε / 10 := sub_nonneg_of_le (div_le_one_of_le₀ (hε₁.trans <| by norm_num) <| by norm_num) have hε₀ : 0 < ε := by sz_positivity calc 4 / 5 * ε = (1 - 1 / 10) * (1 - 9⁻¹) * ε := by norm_num - _ ≤ (1 - ε / 10) * (1 - (↑m)⁻¹) * ((G.nonuniformWitness ε U V).card / U.card) := by + _ ≤ (1 - ε / 10) * (1 - (↑m)⁻¹) * (#(G.nonuniformWitness ε U V) / #U) := by gcongr exacts [mod_cast (show 9 ≤ 100 by norm_num).trans (hundred_le_m hPα hPε hε₁), (le_div_iff₀' <| cast_pos.2 (P.nonempty_of_mem_parts hU).card_pos).2 <| G.le_card_nonuniformWitness hunif] - _ = (1 - ε / 10) * (G.nonuniformWitness ε U V).card * ((1 - (↑m)⁻¹) / U.card) := by + _ = (1 - ε / 10) * #(G.nonuniformWitness ε U V) * ((1 - (↑m)⁻¹) / #U) := by rw [mul_assoc, mul_assoc, mul_div_left_comm] - _ ≤ ((star hP G ε hU V).biUnion id).card * ((1 - (↑m)⁻¹) / U.card) := + _ ≤ #((star hP G ε hU V).biUnion id) * ((1 - (↑m)⁻¹) / #U) := (mul_le_mul_of_nonneg_right (one_sub_eps_mul_card_nonuniformWitness_le_card_star hV hUV hunif hPε hε₁) (by positivity)) - _ ≤ (star hP G ε hU V).card * (m + 1) * ((1 - (↑m)⁻¹) / U.card) := + _ ≤ #(star hP G ε hU V) * (m + 1) * ((1 - (↑m)⁻¹) / #U) := (mul_le_mul_of_nonneg_right card_biUnion_star_le_m_add_one_card_star_mul (by positivity)) - _ ≤ (star hP G ε hU V).card * (m + ↑1) * ((↑1 - (↑m)⁻¹) / (↑4 ^ P.parts.card * m)) := + _ ≤ #(star hP G ε hU V) * (m + ↑1) * ((↑1 - (↑m)⁻¹) / (↑4 ^ #P.parts * m)) := (mul_le_mul_of_nonneg_left (div_le_div_of_nonneg_left hm (by sz_positivity) <| pow_mul_m_le_card_part hP hU) (by positivity)) - _ ≤ (star hP G ε hU V).card / ↑4 ^ P.parts.card := by - rw [mul_assoc, mul_comm ((4 : ℝ) ^ P.parts.card), ← div_div, ← mul_div_assoc, ← mul_comm_div] + _ ≤ #(star hP G ε hU V) / ↑4 ^ #P.parts := by + rw [mul_assoc, mul_comm ((4 : ℝ) ^ #P.parts), ← div_div, ← mul_div_assoc, ← mul_comm_div] refine mul_le_of_le_one_right (by positivity) ?_ have hm : (0 : ℝ) < m := by sz_positivity rw [mul_div_assoc', div_le_one hm, ← one_div, one_sub_div hm.ne', mul_div_assoc', @@ -417,20 +416,20 @@ Those inequalities are the end result of all this hard work. /-- Lower bound on the edge densities between non-uniform parts of `SzemerediRegularity.star`. -/ private theorem edgeDensity_star_not_uniform [Nonempty α] - (hPα : P.parts.card * 16 ^ P.parts.card ≤ card α) (hPε : ↑100 ≤ ↑4 ^ P.parts.card * ε ^ 5) + (hPα : #P.parts * 16 ^ #P.parts ≤ card α) (hPε : ↑100 ≤ ↑4 ^ #P.parts * ε ^ 5) (hε₁ : ε ≤ 1) {hU : U ∈ P.parts} {hV : V ∈ P.parts} (hUVne : U ≠ V) (hUV : ¬G.IsUniform ε U V) : ↑3 / ↑4 * ε ≤ |(∑ ab ∈ (star hP G ε hU V).product (star hP G ε hV U), (G.edgeDensity ab.1 ab.2 : ℝ)) / - ((star hP G ε hU V).card * (star hP G ε hV U).card) - + (#(star hP G ε hU V) * #(star hP G ε hV U)) - (∑ ab ∈ (chunk hP G ε hU).parts.product (chunk hP G ε hV).parts, - (G.edgeDensity ab.1 ab.2 : ℝ)) / (16 : ℝ) ^ P.parts.card| := by + (G.edgeDensity ab.1 ab.2 : ℝ)) / (16 : ℝ) ^ #P.parts| := by rw [show (16 : ℝ) = ↑4 ^ 2 by norm_num, pow_right_comm, sq ((4 : ℝ) ^ _)] set p : ℝ := (∑ ab ∈ (star hP G ε hU V).product (star hP G ε hV U), (G.edgeDensity ab.1 ab.2 : ℝ)) / - ((star hP G ε hU V).card * (star hP G ε hV U).card) + (#(star hP G ε hU V) * #(star hP G ε hV U)) set q : ℝ := (∑ ab ∈ (chunk hP G ε hU).parts.product (chunk hP G ε hV).parts, - (G.edgeDensity ab.1 ab.2 : ℝ)) / (↑4 ^ P.parts.card * ↑4 ^ P.parts.card) + (G.edgeDensity ab.1 ab.2 : ℝ)) / (↑4 ^ #P.parts * ↑4 ^ #P.parts) change _ ≤ |p - q| set r : ℝ := ↑(G.edgeDensity ((star hP G ε hU V).biUnion id) ((star hP G ε hV U).biUnion id)) set s : ℝ := ↑(G.edgeDensity (G.nonuniformWitness ε U V) (G.nonuniformWitness ε V U)) @@ -458,20 +457,20 @@ private theorem edgeDensity_star_not_uniform [Nonempty α] /-- Lower bound on the edge densities between non-uniform parts of `SzemerediRegularity.increment`. -/ -theorem edgeDensity_chunk_not_uniform [Nonempty α] (hPα : P.parts.card * 16 ^ P.parts.card ≤ card α) - (hPε : ↑100 ≤ ↑4 ^ P.parts.card * ε ^ 5) (hε₁ : ε ≤ 1) {hU : U ∈ P.parts} {hV : V ∈ P.parts} +theorem edgeDensity_chunk_not_uniform [Nonempty α] (hPα : #P.parts * 16 ^ #P.parts ≤ card α) + (hPε : ↑100 ≤ ↑4 ^ #P.parts * ε ^ 5) (hε₁ : ε ≤ 1) {hU : U ∈ P.parts} {hV : V ∈ P.parts} (hUVne : U ≠ V) (hUV : ¬G.IsUniform ε U V) : (G.edgeDensity U V : ℝ) ^ 2 - ε ^ 5 / ↑25 + ε ^ 4 / ↑3 ≤ (∑ ab ∈ (chunk hP G ε hU).parts.product (chunk hP G ε hV).parts, - (G.edgeDensity ab.1 ab.2 : ℝ) ^ 2) / ↑16 ^ P.parts.card := + (G.edgeDensity ab.1 ab.2 : ℝ) ^ 2) / ↑16 ^ #P.parts := calc ↑(G.edgeDensity U V) ^ 2 - ε ^ 5 / 25 + ε ^ 4 / ↑3 ≤ ↑(G.edgeDensity U V) ^ 2 - ε ^ 5 / ↑25 + - (star hP G ε hU V).card * (star hP G ε hV U).card / ↑16 ^ P.parts.card * + #(star hP G ε hU V) * #(star hP G ε hV U) / ↑16 ^ #P.parts * (↑9 / ↑16) * ε ^ 2 := by apply add_le_add_left - have Ul : 4 / 5 * ε ≤ (star hP G ε hU V).card / _ := + have Ul : 4 / 5 * ε ≤ #(star hP G ε hU V) / _ := eps_le_card_star_div hPα hPε hε₁ hU hV hUVne hUV - have Vl : 4 / 5 * ε ≤ (star hP G ε hV U).card / _ := + have Vl : 4 / 5 * ε ≤ #(star hP G ε hV U) / _ := eps_le_card_star_div hPα hPε hε₁ hV hU hUVne.symm fun h => hUV h.symm rw [show (16 : ℝ) = ↑4 ^ 2 by norm_num, pow_right_comm, sq ((4 : ℝ) ^ _), ← _root_.div_mul_div_comm, mul_assoc] @@ -487,7 +486,7 @@ theorem edgeDensity_chunk_not_uniform [Nonempty α] (hPα : P.parts.card * 16 ^ · norm_num positivity _ ≤ (∑ ab ∈ (chunk hP G ε hU).parts.product (chunk hP G ε hV).parts, - (G.edgeDensity ab.1 ab.2 : ℝ) ^ 2) / ↑16 ^ P.parts.card := by + (G.edgeDensity ab.1 ab.2 : ℝ) ^ 2) / ↑16 ^ #P.parts := by have t : (star hP G ε hU V).product (star hP G ε hV U) ⊆ (chunk hP G ε hU).parts.product (chunk hP G ε hV).parts := product_subset_product star_subset_chunk star_subset_chunk @@ -510,14 +509,13 @@ theorem edgeDensity_chunk_not_uniform [Nonempty α] (hPα : P.parts.card * 16 ^ /-- Lower bound on the edge densities between parts of `SzemerediRegularity.increment`. This is the blanket lower bound used the uniform parts. -/ -theorem edgeDensity_chunk_uniform [Nonempty α] (hPα : P.parts.card * 16 ^ P.parts.card ≤ card α) - (hPε : ↑100 ≤ ↑4 ^ P.parts.card * ε ^ 5) (hU : U ∈ P.parts) (hV : V ∈ P.parts) : +theorem edgeDensity_chunk_uniform [Nonempty α] (hPα : #P.parts * 16 ^ #P.parts ≤ card α) + (hPε : ↑100 ≤ ↑4 ^ #P.parts * ε ^ 5) (hU : U ∈ P.parts) (hV : V ∈ P.parts) : (G.edgeDensity U V : ℝ) ^ 2 - ε ^ 5 / ↑25 ≤ (∑ ab ∈ (chunk hP G ε hU).parts.product (chunk hP G ε hV).parts, - (G.edgeDensity ab.1 ab.2 : ℝ) ^ 2) / ↑16 ^ P.parts.card := by + (G.edgeDensity ab.1 ab.2 : ℝ) ^ 2) / ↑16 ^ #P.parts := by apply (edgeDensity_chunk_aux (hP := hP) hPα hPε hU hV).trans - have key : ↑16 ^ P.parts.card = - (((chunk hP G ε hU).parts ×ˢ (chunk hP G ε hV).parts).card : ℝ) := by + have key : (16 : ℝ) ^ #P.parts = #((chunk hP G ε hU).parts ×ˢ (chunk hP G ε hV).parts) := by rw [card_product, cast_mul, card_chunk (m_pos hPα).ne', card_chunk (m_pos hPα).ne', ← cast_mul, ← mul_pow]; norm_cast simp_rw [key] diff --git a/Mathlib/Combinatorics/SimpleGraph/Regularity/Energy.lean b/Mathlib/Combinatorics/SimpleGraph/Regularity/Energy.lean index ad2bb21be97a3..2784e3dcc622c 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Regularity/Energy.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Regularity/Energy.lean @@ -33,7 +33,7 @@ namespace Finpartition /-- The energy of a partition, also known as index. Auxiliary quantity for Szemerédi's regularity lemma. -/ def energy : ℚ := - ((∑ uv ∈ P.parts.offDiag, G.edgeDensity uv.1 uv.2 ^ 2) : ℚ) / (P.parts.card : ℚ) ^ 2 + ((∑ uv ∈ P.parts.offDiag, G.edgeDensity uv.1 uv.2 ^ 2) : ℚ) / (#P.parts : ℚ) ^ 2 theorem energy_nonneg : 0 ≤ P.energy G := by exact div_nonneg (Finset.sum_nonneg fun _ _ => sq_nonneg _) <| sq_nonneg _ @@ -41,10 +41,10 @@ theorem energy_nonneg : 0 ≤ P.energy G := by theorem energy_le_one : P.energy G ≤ 1 := div_le_of_le_mul₀ (sq_nonneg _) zero_le_one <| calc - ∑ uv ∈ P.parts.offDiag, G.edgeDensity uv.1 uv.2 ^ 2 ≤ P.parts.offDiag.card • (1 : ℚ) := + ∑ uv ∈ P.parts.offDiag, G.edgeDensity uv.1 uv.2 ^ 2 ≤ #P.parts.offDiag • (1 : ℚ) := sum_le_card_nsmul _ _ 1 fun _ _ => (sq_le_one_iff <| G.edgeDensity_nonneg _ _).2 <| G.edgeDensity_le_one _ _ - _ = P.parts.offDiag.card := Nat.smul_one_eq_cast _ + _ = #P.parts.offDiag := Nat.smul_one_eq_cast _ _ ≤ _ := by rw [offDiag_card, one_mul] norm_cast @@ -53,7 +53,7 @@ theorem energy_le_one : P.energy G ≤ 1 := @[simp, norm_cast] theorem coe_energy {𝕜 : Type*} [LinearOrderedField 𝕜] : (P.energy G : 𝕜) = - (∑ uv ∈ P.parts.offDiag, (G.edgeDensity uv.1 uv.2 : 𝕜) ^ 2) / (P.parts.card : 𝕜) ^ 2 := by + (∑ uv ∈ P.parts.offDiag, (G.edgeDensity uv.1 uv.2 : 𝕜) ^ 2) / (#P.parts : 𝕜) ^ 2 := by rw [energy]; norm_cast end Finpartition diff --git a/Mathlib/Combinatorics/SimpleGraph/Regularity/Equitabilise.lean b/Mathlib/Combinatorics/SimpleGraph/Regularity/Equitabilise.lean index 96d140587f9cc..a5b1a55dfb92a 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Regularity/Equitabilise.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Regularity/Equitabilise.lean @@ -32,16 +32,16 @@ namespace Finpartition variable {α : Type*} [DecidableEq α] {s t : Finset α} {m n a b : ℕ} {P : Finpartition s} -/-- Given a partition `P` of `s`, as well as a proof that `a * m + b * (m + 1) = s.card`, we can +/-- Given a partition `P` of `s`, as well as a proof that `a * m + b * (m + 1) = #s`, 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 (provided `m > 0`, because a partition does not have parts of size `0`) there are `a` parts of size `m` and hence `a + b` parts in total. -/ -theorem equitabilise_aux (hs : a * m + b * (m + 1) = s.card) : +theorem equitabilise_aux (hs : a * m + b * (m + 1) = #s) : ∃ Q : Finpartition s, - (∀ x : Finset α, x ∈ Q.parts → x.card = m ∨ x.card = m + 1) ∧ - (∀ x, x ∈ P.parts → (x \ (Q.parts.filter fun y => y ⊆ x).biUnion id).card ≤ m) ∧ - (Q.parts.filter fun i => card i = m + 1).card = b := by + (∀ x : Finset α, x ∈ Q.parts → #x = m ∨ #x = m + 1) ∧ + (∀ x, x ∈ P.parts → #(x \ {y ∈ Q.parts | y ⊆ x}.biUnion id) ≤ m) ∧ + #{i ∈ Q.parts | #i = m + 1} = b := by -- Get rid of the easy case `m = 0` obtain rfl | m_pos := m.eq_zero_or_pos · refine ⟨⊥, by simp, ?_, by simpa [Finset.filter_true_of_mem] using hs.symm⟩ @@ -63,7 +63,7 @@ theorem equitabilise_aux (hs : a * m + b * (m + 1) = s.card) : set n := if 0 < a then m else m + 1 with hn -- Some easy facts about it obtain ⟨hn₀, hn₁, hn₂, hn₃⟩ : 0 < n ∧ n ≤ m + 1 ∧ n ≤ a * m + b * (m + 1) ∧ - ite (0 < a) (a - 1) a * m + ite (0 < a) b (b - 1) * (m + 1) = s.card - n := by + ite (0 < a) (a - 1) a * m + ite (0 < a) b (b - 1) * (m + 1) = #s - n := by rw [hn, ← hs] split_ifs with h <;> rw [tsub_mul, one_mul] · refine ⟨m_pos, le_succ _, le_add_right (Nat.le_mul_of_pos_left _ ‹0 < a›), ?_⟩ @@ -77,10 +77,10 @@ theorem equitabilise_aux (hs : a * m + b * (m + 1) = s.card) : least one part `u` of `P` has size `m + 1` (in which case we take `t` to be an arbitrary subset of `u` of size `n`). The rest of each branch is just tedious calculations to satisfy the induction hypothesis. -/ - by_cases h : ∀ u ∈ P.parts, card u < m + 1 + by_cases h : ∀ u ∈ P.parts, #u < m + 1 · obtain ⟨t, hts, htn⟩ := exists_subset_card_eq (hn₂.trans_eq hs) have ht : t.Nonempty := by rwa [← card_pos, htn] - have hcard : ite (0 < a) (a - 1) a * m + ite (0 < a) b (b - 1) * (m + 1) = (s \ t).card := by + have hcard : ite (0 < a) (a - 1) a * m + ite (0 < a) b (b - 1) * (m + 1) = #(s \ t) := by rw [card_sdiff ‹t ⊆ s›, htn, hn₃] obtain ⟨R, hR₁, _, hR₃⟩ := @ih (s \ t) (sdiff_ssubset hts ‹t.Nonempty›) (if 0 < a then a - 1 else a) @@ -99,7 +99,7 @@ theorem equitabilise_aux (hs : a * m + b * (m + 1) = s.card) : obtain ⟨u, hu₁, hu₂⟩ := h obtain ⟨t, htu, htn⟩ := exists_subset_card_eq (hn₁.trans hu₂) have ht : t.Nonempty := by rwa [← card_pos, htn] - have hcard : ite (0 < a) (a - 1) a * m + ite (0 < a) b (b - 1) * (m + 1) = (s \ t).card := by + have hcard : ite (0 < a) (a - 1) a * m + ite (0 < a) b (b - 1) * (m + 1) = #(s \ t) := by rw [card_sdiff (htu.trans <| P.le hu₁), htn, hn₃] obtain ⟨R, hR₁, hR₂, hR₃⟩ := @ih (s \ t) (sdiff_ssubset (htu.trans <| P.le hu₁) ht) (if 0 < a then a - 1 else a) @@ -136,9 +136,9 @@ theorem equitabilise_aux (hs : a * m + b * (m + 1) = s.card) : · rw [card_insert_of_not_mem, hR₃, if_neg h, Nat.sub_add_cancel (hab.resolve_left h)] intro H; exact ht.ne_empty (le_sdiff_iff.1 <| R.le <| filter_subset _ _ H) -variable (h : a * m + b * (m + 1) = s.card) +variable (h : a * m + b * (m + 1) = #s) -/-- Given a partition `P` of `s`, as well as a proof that `a * m + b * (m + 1) = s.card`, build a +/-- Given a partition `P` of `s`, as well as a proof that `a * m + b * (m + 1) = #s`, build 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 (provided `m > 0`, because a partition does not have parts of size `0`) there are `a` parts of size `m` and @@ -149,7 +149,7 @@ noncomputable def equitabilise : Finpartition s := variable {h} theorem card_eq_of_mem_parts_equitabilise : - t ∈ (P.equitabilise h).parts → t.card = m ∨ t.card = m + 1 := + t ∈ (P.equitabilise h).parts → #t = m ∨ #t = m + 1 := (P.equitabilise_aux h).choose_spec.1 _ theorem equitabilise_isEquipartition : (P.equitabilise h).IsEquipartition := @@ -157,18 +157,16 @@ theorem equitabilise_isEquipartition : (P.equitabilise h).IsEquipartition := variable (P h) -theorem card_filter_equitabilise_big : - ((P.equitabilise h).parts.filter fun u : Finset α => u.card = m + 1).card = b := +theorem card_filter_equitabilise_big : #{u ∈ (P.equitabilise h).parts | #u = m + 1} = b := (P.equitabilise_aux h).choose_spec.2.2 theorem card_filter_equitabilise_small (hm : m ≠ 0) : - ((P.equitabilise h).parts.filter fun u : Finset α => u.card = m).card = a := by + #{u ∈ (P.equitabilise h).parts | #u = m} = a := by refine (mul_eq_mul_right_iff.1 <| (add_left_inj (b * (m + 1))).1 ?_).resolve_right hm rw [h, ← (P.equitabilise h).sum_card_parts] have hunion : (P.equitabilise h).parts = - ((P.equitabilise h).parts.filter fun u => u.card = m) ∪ - (P.equitabilise h).parts.filter fun u => u.card = m + 1 := by + {u ∈ (P.equitabilise h).parts | #u = m} ∪ {u ∈ (P.equitabilise h).parts | #u = m + 1} := by rw [← filter_or, filter_true_of_mem] exact fun x => card_eq_of_mem_parts_equitabilise nth_rw 2 [hunion] @@ -179,23 +177,23 @@ theorem card_filter_equitabilise_small (hm : m ≠ 0) : apply succ_ne_self m _ exact (hb i h).symm.trans (ha i h) -theorem card_parts_equitabilise (hm : m ≠ 0) : (P.equitabilise h).parts.card = a + b := by +theorem card_parts_equitabilise (hm : m ≠ 0) : #(P.equitabilise h).parts = a + b := by rw [← filter_true_of_mem fun x => card_eq_of_mem_parts_equitabilise, filter_or, card_union_of_disjoint, P.card_filter_equitabilise_small _ hm, P.card_filter_equitabilise_big] -- Porting note (#11187): was `infer_instance` exact disjoint_filter.2 fun x _ h₀ h₁ => Nat.succ_ne_self m <| h₁.symm.trans h₀ theorem card_parts_equitabilise_subset_le : - t ∈ P.parts → (t \ ((P.equitabilise h).parts.filter fun u => u ⊆ t).biUnion id).card ≤ m := + t ∈ P.parts → #(t \ {u ∈ (P.equitabilise h).parts | u ⊆ t}.biUnion id) ≤ m := (Classical.choose_spec <| P.equitabilise_aux h).2.1 t variable (s) /-- We can find equipartitions of arbitrary size. -/ -theorem exists_equipartition_card_eq (hn : n ≠ 0) (hs : n ≤ s.card) : - ∃ P : Finpartition s, P.IsEquipartition ∧ P.parts.card = n := by +theorem exists_equipartition_card_eq (hn : n ≠ 0) (hs : n ≤ #s) : + ∃ P : Finpartition s, P.IsEquipartition ∧ #P.parts = n := by rw [← pos_iff_ne_zero] at hn - have : (n - s.card % n) * (s.card / n) + s.card % n * (s.card / n + 1) = s.card := by + have : (n - #s % n) * (#s / n) + #s % n * (#s / n + 1) = #s := by rw [tsub_mul, mul_add, ← add_assoc, tsub_add_cancel_of_le (Nat.mul_le_mul_right _ (mod_lt _ hn).le), mul_one, add_comm, mod_add_div] diff --git a/Mathlib/Combinatorics/SimpleGraph/Regularity/Increment.lean b/Mathlib/Combinatorics/SimpleGraph/Regularity/Increment.lean index a6ff90799f9cb..6f97ef8109a10 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Regularity/Increment.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Regularity/Increment.lean @@ -41,7 +41,7 @@ open scoped SzemerediRegularity.Positivity variable {α : Type*} [Fintype α] [DecidableEq α] {P : Finpartition (univ : Finset α)} (hP : P.IsEquipartition) (G : SimpleGraph α) [DecidableRel G.Adj] (ε : ℝ) -local notation3 "m" => (card α / stepBound P.parts.card : ℕ) +local notation3 "m" => (card α / stepBound #P.parts : ℕ) namespace SzemerediRegularity @@ -59,11 +59,11 @@ open Finpartition Finpartition.IsEquipartition variable {hP G ε} /-- The increment partition has a prescribed (very big) size in terms of the original partition. -/ -theorem card_increment (hPα : P.parts.card * 16 ^ P.parts.card ≤ card α) (hPG : ¬P.IsUniform G ε) : - (increment hP G ε).parts.card = stepBound P.parts.card := by - have hPα' : stepBound P.parts.card ≤ card α := +theorem card_increment (hPα : #P.parts * 16 ^ #P.parts ≤ card α) (hPG : ¬P.IsUniform G ε) : + #(increment hP G ε).parts = stepBound #P.parts := by + have hPα' : stepBound #P.parts ≤ card α := (mul_le_mul_left' (pow_le_pow_left' (by norm_num) _) _).trans hPα - have hPpos : 0 < stepBound P.parts.card := stepBound_pos (nonempty_of_not_uniform hPG).card_pos + have hPpos : 0 < stepBound #P.parts := stepBound_pos (nonempty_of_not_uniform hPG).card_pos rw [increment, card_bind] simp_rw [chunk, apply_dite Finpartition.parts, apply_dite card, sum_dite] rw [sum_const_nat, sum_const_nat, univ_eq_attach, univ_eq_attach, card_attach, card_attach] @@ -119,10 +119,10 @@ private lemma pairwiseDisjoint_distinctPairs : variable [Nonempty α] lemma le_sum_distinctPairs_edgeDensity_sq (x : {i // i ∈ P.parts.offDiag}) (hε₁ : ε ≤ 1) - (hPα : P.parts.card * 16 ^ P.parts.card ≤ card α) (hPε : ↑100 ≤ ↑4 ^ P.parts.card * ε ^ 5) : + (hPα : #P.parts * 16 ^ #P.parts ≤ card α) (hPε : ↑100 ≤ ↑4 ^ #P.parts * ε ^ 5) : (G.edgeDensity x.1.1 x.1.2 : ℝ) ^ 2 + ((if G.IsUniform ε x.1.1 x.1.2 then 0 else ε ^ 4 / 3) - ε ^ 5 / 25) ≤ - (∑ i ∈ distinctPairs hP G ε x, G.edgeDensity i.1 i.2 ^ 2 : ℝ) / 16 ^ P.parts.card := by + (∑ i ∈ distinctPairs hP G ε x, G.edgeDensity i.1 i.2 ^ 2 : ℝ) / 16 ^ #P.parts := by rw [distinctPairs, ← add_sub_assoc, add_sub_right_comm] split_ifs with h · rw [add_zero] @@ -130,18 +130,18 @@ lemma le_sum_distinctPairs_edgeDensity_sq (x : {i // i ∈ P.parts.offDiag}) (h · exact edgeDensity_chunk_not_uniform hPα hPε hε₁ (mem_offDiag.1 x.2).2.2 h /-- The increment partition has energy greater than the original one by a known fixed amount. -/ -theorem energy_increment (hP : P.IsEquipartition) (hP₇ : 7 ≤ P.parts.card) - (hPε : 100 ≤ 4 ^ P.parts.card * ε ^ 5) (hPα : P.parts.card * 16 ^ P.parts.card ≤ card α) +theorem energy_increment (hP : P.IsEquipartition) (hP₇ : 7 ≤ #P.parts) + (hPε : 100 ≤ 4 ^ #P.parts * ε ^ 5) (hPα : #P.parts * 16 ^ #P.parts ≤ card α) (hPG : ¬P.IsUniform G ε) (hε₀ : 0 ≤ ε) (hε₁ : ε ≤ 1) : ↑(P.energy G) + ε ^ 5 / 4 ≤ (increment hP G ε).energy G := by calc _ = (∑ x ∈ P.parts.offDiag, (G.edgeDensity x.1 x.2 : ℝ) ^ 2 + - P.parts.card ^ 2 * (ε ^ 5 / 4) : ℝ) / P.parts.card ^ 2 := by + #P.parts ^ 2 * (ε ^ 5 / 4) : ℝ) / #P.parts ^ 2 := by rw [coe_energy, add_div, mul_div_cancel_left₀]; positivity _ ≤ (∑ x ∈ P.parts.offDiag.attach, (∑ i ∈ distinctPairs hP G ε x, - G.edgeDensity i.1 i.2 ^ 2 : ℝ) / 16 ^ P.parts.card) / P.parts.card ^ 2 := ?_ + G.edgeDensity i.1 i.2 ^ 2 : ℝ) / 16 ^ #P.parts) / #P.parts ^ 2 := ?_ _ = (∑ x ∈ P.parts.offDiag.attach, ∑ i ∈ distinctPairs hP G ε x, - G.edgeDensity i.1 i.2 ^ 2 : ℝ) / (increment hP G ε).parts.card ^ 2 := by + G.edgeDensity i.1 i.2 ^ 2 : ℝ) / #(increment hP G ε).parts ^ 2 := by rw [card_increment hPα hPG, coe_stepBound, mul_pow, pow_right_comm, div_mul_eq_div_div_swap, ← sum_div]; norm_num _ ≤ _ := by @@ -153,7 +153,7 @@ theorem energy_increment (hP : P.IsEquipartition) (hP₇ : 7 ≤ P.parts.card) rw [Finpartition.IsUniform, not_le, mul_tsub, mul_one, ← offDiag_card] at hPG calc _ ≤ ∑ x ∈ P.parts.offDiag, (edgeDensity G x.1 x.2 : ℝ) ^ 2 + - ((nonUniforms P G ε).card * (ε ^ 4 / 3) - P.parts.offDiag.card * (ε ^ 5 / 25)) := + (#(nonUniforms P G ε) * (ε ^ 4 / 3) - #P.parts.offDiag * (ε ^ 5 / 25)) := add_le_add_left ?_ _ _ = ∑ x ∈ P.parts.offDiag, ((G.edgeDensity x.1 x.2 : ℝ) ^ 2 + ((if G.IsUniform ε x.1 x.2 then (0 : ℝ) else ε ^ 4 / 3) - ε ^ 5 / 25) : ℝ) := by @@ -163,8 +163,8 @@ theorem energy_increment (hP : P.IsEquipartition) (hP₇ : 7 ≤ P.parts.card) _ = _ := (sum_attach ..).symm _ ≤ _ := sum_le_sum fun i _ ↦ le_sum_distinctPairs_edgeDensity_sq i hε₁ hPα hPε calc - _ = (6/7 * P.parts.card ^ 2) * ε ^ 5 * (7 / 24) := by ring - _ ≤ P.parts.offDiag.card * ε ^ 5 * (22 / 75) := by + _ = (6/7 * #P.parts ^ 2) * ε ^ 5 * (7 / 24) := by ring + _ ≤ #P.parts.offDiag * ε ^ 5 * (22 / 75) := by gcongr ?_ * _ * ?_ · rw [← mul_div_right_comm, div_le_iff₀ (by norm_num), offDiag_card] norm_cast @@ -172,7 +172,7 @@ theorem energy_increment (hP : P.IsEquipartition) (hP₇ : 7 ≤ P.parts.card) refine le_tsub_of_add_le_left ?_ nlinarith · norm_num - _ = (P.parts.offDiag.card * ε * (ε ^ 4 / 3) - P.parts.offDiag.card * (ε ^ 5 / 25)) := by ring - _ ≤ ((nonUniforms P G ε).card * (ε ^ 4 / 3) - P.parts.offDiag.card * (ε ^ 5 / 25)) := by gcongr + _ = (#P.parts.offDiag * ε * (ε ^ 4 / 3) - #P.parts.offDiag * (ε ^ 5 / 25)) := by ring + _ ≤ (#(nonUniforms P G ε) * (ε ^ 4 / 3) - #P.parts.offDiag * (ε ^ 5 / 25)) := by gcongr end SzemerediRegularity diff --git a/Mathlib/Combinatorics/SimpleGraph/Regularity/Lemma.lean b/Mathlib/Combinatorics/SimpleGraph/Regularity/Lemma.lean index bee7ea6b54e48..dd9d05a8d0c75 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Regularity/Lemma.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Regularity/Lemma.lean @@ -71,7 +71,7 @@ variable {α : Type*} [DecidableEq α] [Fintype α] (G : SimpleGraph α) [Decida `ε`-uniform equipartition of bounded size (where the bound does not depend on the graph). -/ theorem szemeredi_regularity (hε : 0 < ε) (hl : l ≤ card α) : ∃ P : Finpartition univ, - P.IsEquipartition ∧ l ≤ P.parts.card ∧ P.parts.card ≤ bound ε l ∧ P.IsUniform G ε := by + P.IsEquipartition ∧ l ≤ #P.parts ∧ #P.parts ≤ bound ε l ∧ P.IsUniform G ε := by obtain hα | hα := le_total (card α) (bound ε l) -- If `card α ≤ bound ε l`, then the partition into singletons is acceptable. · refine ⟨⊥, bot_isEquipartition _, ?_⟩ @@ -79,7 +79,7 @@ theorem szemeredi_regularity (hε : 0 < ε) (hl : l ≤ card α) : exact ⟨hl, hα, bot_isUniform _ hε⟩ -- Else, let's start from a dummy equipartition of size `initialBound ε l`. let t := initialBound ε l - have htα : t ≤ (univ : Finset α).card := + have htα : t ≤ #(univ : Finset α) := (initialBound_le_bound _ _).trans (by rwa [Finset.card_univ]) obtain ⟨dum, hdum₁, hdum₂⟩ := exists_equipartition_card_eq (univ : Finset α) (initialBound_pos _ _).ne' htα @@ -93,8 +93,8 @@ theorem szemeredi_regularity (hε : 0 < ε) (hl : l ≤ card α) : have : Nonempty α := by rw [← Fintype.card_pos_iff] exact (bound_pos _ _).trans_le hα - suffices h : ∀ i, ∃ P : Finpartition (univ : Finset α), P.IsEquipartition ∧ t ≤ P.parts.card ∧ - P.parts.card ≤ stepBound^[i] t ∧ (P.IsUniform G ε ∨ ε ^ 5 / 4 * i ≤ P.energy G) by + suffices h : ∀ i, ∃ P : Finpartition (univ : Finset α), P.IsEquipartition ∧ t ≤ #P.parts ∧ + #P.parts ≤ stepBound^[i] t ∧ (P.IsUniform G ε ∨ ε ^ 5 / 4 * i ≤ P.energy G) by -- For `i > 4 / ε ^ 5` we know that the partition we get can't have energy `≥ ε ^ 5 / 4 * i > 1`, -- so it must instead be `ε`-uniform and we won. obtain ⟨P, hP₁, hP₂, hP₃, hP₄⟩ := h (⌊4 / ε ^ 5⌋₊ + 1) @@ -126,7 +126,7 @@ theorem szemeredi_regularity (hε : 0 < ε) (hl : l ≤ card α) : -- Else, `P` must instead have energy at least `ε ^ 5 / 4 * i`. replace hP₄ := hP₄.resolve_left huniform -- We gather a few numerical facts. - have hεl' : 100 ≤ 4 ^ P.parts.card * ε ^ 5 := + have hεl' : 100 ≤ 4 ^ #P.parts * ε ^ 5 := (hundred_lt_pow_initialBound_mul hε l).le.trans (mul_le_mul_of_nonneg_right (pow_right_mono₀ (by norm_num) hP₂) <| by positivity) have hi : (i : ℝ) ≤ 4 / ε ^ 5 := by @@ -134,9 +134,9 @@ theorem szemeredi_regularity (hε : 0 < ε) (hl : l ≤ card α) : rw [div_mul_eq_mul_div, div_le_iff₀ (show (0 : ℝ) < 4 by norm_num)] at hi norm_num at hi rwa [le_div_iff₀' (pow_pos hε _)] - have hsize : P.parts.card ≤ stepBound^[⌊4 / ε ^ 5⌋₊] t := + have hsize : #P.parts ≤ stepBound^[⌊4 / ε ^ 5⌋₊] t := hP₃.trans (monotone_iterate_of_id_le le_stepBound (Nat.le_floor hi) _) - have hPα : P.parts.card * 16 ^ P.parts.card ≤ card α := + have hPα : #P.parts * 16 ^ #P.parts ≤ card α := (Nat.mul_le_mul hsize (Nat.pow_le_pow_of_le_right (by norm_num) hsize)).trans hα -- We return the increment equipartition of `P`, which has energy `≥ ε ^ 5 / 4 * (i + 1)`. refine ⟨increment hP₁ G ε, increment_isEquipartition hP₁ G ε, ?_, ?_, Or.inr <| le_trans ?_ <| diff --git a/Mathlib/Combinatorics/SimpleGraph/Regularity/Uniform.lean b/Mathlib/Combinatorics/SimpleGraph/Regularity/Uniform.lean index 572c8e65f679c..6d9c7efcbf156 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Regularity/Uniform.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Regularity/Uniform.lean @@ -55,8 +55,8 @@ variable (G : SimpleGraph α) [DecidableRel G.Adj] (ε : 𝕜) {s t : Finset α} to the density of any big enough pair of subsets. Intuitively, the edges between them are random-like. -/ def IsUniform (s t : Finset α) : Prop := - ∀ ⦃s'⦄, s' ⊆ s → ∀ ⦃t'⦄, t' ⊆ t → (s.card : 𝕜) * ε ≤ s'.card → - (t.card : 𝕜) * ε ≤ t'.card → |(G.edgeDensity s' t' : 𝕜) - (G.edgeDensity s t : 𝕜)| < ε + ∀ ⦃s'⦄, s' ⊆ s → ∀ ⦃t'⦄, t' ⊆ t → (#s : 𝕜) * ε ≤ #s' → + (#t : 𝕜) * ε ≤ #t' → |(G.edgeDensity s' t' : 𝕜) - (G.edgeDensity s t : 𝕜)| < ε variable {G ε} @@ -105,8 +105,8 @@ theorem not_isUniform_zero : ¬G.IsUniform (0 : 𝕜) s t := fun h => (abs_nonneg _).not_lt <| h (empty_subset _) (empty_subset _) (by simp) (by simp) theorem not_isUniform_iff : - ¬G.IsUniform ε s t ↔ ∃ s', s' ⊆ s ∧ ∃ t', t' ⊆ t ∧ ↑s.card * ε ≤ s'.card ∧ - ↑t.card * ε ≤ t'.card ∧ ε ≤ |G.edgeDensity s' t' - G.edgeDensity s t| := by + ¬G.IsUniform ε s t ↔ ∃ s', s' ⊆ s ∧ ∃ t', t' ⊆ t ∧ #s * ε ≤ #s' ∧ + #t * ε ≤ #t' ∧ ε ≤ |G.edgeDensity s' t' - G.edgeDensity s t| := by unfold IsUniform simp only [not_forall, not_lt, exists_prop, exists_and_left, Rat.cast_abs, Rat.cast_sub] @@ -128,7 +128,7 @@ theorem left_nonuniformWitnesses_subset (h : ¬G.IsUniform ε s t) : exact (not_isUniform_iff.1 h).choose_spec.1 theorem left_nonuniformWitnesses_card (h : ¬G.IsUniform ε s t) : - (s.card : 𝕜) * ε ≤ (G.nonuniformWitnesses ε s t).1.card := by + #s * ε ≤ #(G.nonuniformWitnesses ε s t).1 := by rw [nonuniformWitnesses, dif_pos h] exact (not_isUniform_iff.1 h).choose_spec.2.choose_spec.2.1 @@ -138,7 +138,7 @@ theorem right_nonuniformWitnesses_subset (h : ¬G.IsUniform ε s t) : exact (not_isUniform_iff.1 h).choose_spec.2.choose_spec.1 theorem right_nonuniformWitnesses_card (h : ¬G.IsUniform ε s t) : - (t.card : 𝕜) * ε ≤ (G.nonuniformWitnesses ε s t).2.card := by + #t * ε ≤ #(G.nonuniformWitnesses ε s t).2 := by rw [nonuniformWitnesses, dif_pos h] exact (not_isUniform_iff.1 h).choose_spec.2.choose_spec.2.2.1 @@ -162,7 +162,7 @@ theorem nonuniformWitness_subset (h : ¬G.IsUniform ε s t) : G.nonuniformWitnes · exact G.right_nonuniformWitnesses_subset fun i => h i.symm theorem le_card_nonuniformWitness (h : ¬G.IsUniform ε s t) : - (s.card : 𝕜) * ε ≤ (G.nonuniformWitness ε s t).card := by + #s * ε ≤ #(G.nonuniformWitness ε s t) := by unfold nonuniformWitness split_ifs · exact G.left_nonuniformWitnesses_card h @@ -224,7 +224,7 @@ theorem nonUniforms_bot (hε : 0 < ε) : (⊥ : Finpartition A).nonUniforms G ε /-- A finpartition of a graph's vertex set is `ε`-uniform (aka `ε`-regular) iff the proportion of its pairs of parts that are not `ε`-uniform is at most `ε`. -/ def IsUniform (ε : 𝕜) : Prop := - ((P.nonUniforms G ε).card : 𝕜) ≤ (P.parts.card * (P.parts.card - 1) : ℕ) * ε + (#(P.nonUniforms G ε) : 𝕜) ≤ (#P.parts * (#P.parts - 1) : ℕ) * ε lemma bot_isUniform (hε : 0 < ε) : (⊥ : Finpartition A).IsUniform G ε := by rw [Finpartition.IsUniform, Finpartition.card_bot, nonUniforms_bot _ hε, Finset.card_empty, @@ -252,7 +252,7 @@ variable (P G ε) (s : Finset α) /-- A choice of witnesses of non-uniformity among the parts of a finpartition. -/ noncomputable def nonuniformWitnesses : Finset (Finset α) := - (P.parts.filter fun t => s ≠ t ∧ ¬G.IsUniform ε s t).image (G.nonuniformWitness ε s) + {t ∈ P.parts | s ≠ t ∧ ¬G.IsUniform ε s t}.image (G.nonuniformWitness ε s) variable {P G ε s} {t : Finset α} @@ -265,11 +265,10 @@ theorem nonuniformWitness_mem_nonuniformWitnesses (h : ¬G.IsUniform ε s t) (ht open SimpleGraph in lemma IsEquipartition.card_interedges_sparsePairs_le' (hP : P.IsEquipartition) (hε : 0 ≤ ε) : - ((P.sparsePairs G ε).biUnion fun (U, V) ↦ G.interedges U V).card ≤ - ε * (A.card + P.parts.card) ^ 2 := by + #((P.sparsePairs G ε).biUnion fun (U, V) ↦ G.interedges U V) ≤ ε * (#A + #P.parts) ^ 2 := by calc - _ ≤ ∑ UV ∈ P.sparsePairs G ε, ((G.interedges UV.1 UV.2).card : 𝕜) := mod_cast card_biUnion_le - _ ≤ ∑ UV ∈ P.sparsePairs G ε, ε * (UV.1.card * UV.2.card) := ?_ + _ ≤ ∑ UV ∈ P.sparsePairs G ε, (#(G.interedges UV.1 UV.2) : 𝕜) := mod_cast card_biUnion_le + _ ≤ ∑ UV ∈ P.sparsePairs G ε, ε * (#UV.1 * #UV.2) := ?_ _ ≤ _ := sum_le_sum_of_subset_of_nonneg (filter_subset _ _) fun i _ _ ↦ by positivity _ = _ := (mul_sum _ _ _).symm _ ≤ _ := mul_le_mul_of_nonneg_left ?_ hε @@ -281,9 +280,9 @@ lemma IsEquipartition.card_interedges_sparsePairs_le' (hP : P.IsEquipartition) (Nat.cast_pos.2 (P.nonempty_of_mem_parts hUV.2.1).card_pos) norm_cast calc - (_ : ℕ) ≤ _ := sum_le_card_nsmul P.parts.offDiag (fun i ↦ i.1.card * i.2.card) - ((A.card / P.parts.card + 1)^2 : ℕ) ?_ - _ ≤ (P.parts.card * (A.card / P.parts.card) + P.parts.card) ^ 2 := ?_ + (_ : ℕ) ≤ _ := sum_le_card_nsmul P.parts.offDiag (fun i ↦ #i.1 * #i.2) + ((#A / #P.parts + 1)^2 : ℕ) ?_ + _ ≤ (#P.parts * (#A / #P.parts) + #P.parts) ^ 2 := ?_ _ ≤ _ := Nat.pow_le_pow_of_le_left (add_le_add_right (Nat.mul_div_le _ _) _) _ · simp only [Prod.forall, Finpartition.mk_mem_nonUniforms, and_imp, mem_offDiag, sq] rintro U V hU hV - @@ -293,10 +292,10 @@ lemma IsEquipartition.card_interedges_sparsePairs_le' (hP : P.IsEquipartition) exact Nat.sub_le _ _ lemma IsEquipartition.card_interedges_sparsePairs_le (hP : P.IsEquipartition) (hε : 0 ≤ ε) : - ((P.sparsePairs G ε).biUnion fun (U, V) ↦ G.interedges U V).card ≤ 4 * ε * A.card ^ 2 := by + #((P.sparsePairs G ε).biUnion fun (U, V) ↦ G.interedges U V) ≤ 4 * ε * #A ^ 2 := by calc _ ≤ _ := hP.card_interedges_sparsePairs_le' hε - _ ≤ ε * (A.card + A.card)^2 := by gcongr; exact P.card_parts_le_card + _ ≤ ε * (#A + #A)^2 := by gcongr; exact P.card_parts_le_card _ = _ := by ring private lemma aux {i j : ℕ} (hj : 0 < j) : j * (j - 1) * (i / j + 1) ^ 2 < (i + j) ^ 2 := by @@ -307,44 +306,44 @@ private lemma aux {i j : ℕ} (hj : 0 < j) : j * (j - 1) * (i / j + 1) ^ 2 < (i exact Nat.pow_le_pow_of_le_left (add_le_add_right (Nat.mul_div_le i j) _) _ lemma IsEquipartition.card_biUnion_offDiag_le' (hP : P.IsEquipartition) : - ((P.parts.biUnion offDiag).card : 𝕜) ≤ A.card * (A.card + P.parts.card) / P.parts.card := by + (#(P.parts.biUnion offDiag) : 𝕜) ≤ #A * (#A + #P.parts) / #P.parts := by obtain h | h := P.parts.eq_empty_or_nonempty · simp [h] calc - _ ≤ (P.parts.card : 𝕜) * (↑(A.card / P.parts.card) * ↑(A.card / P.parts.card + 1)) := + _ ≤ (#P.parts : 𝕜) * (↑(#A / #P.parts) * ↑(#A / #P.parts + 1)) := mod_cast card_biUnion_le_card_mul _ _ _ fun U hU ↦ ?_ - _ = P.parts.card * ↑(A.card / P.parts.card) * ↑(A.card / P.parts.card + 1) := by rw [mul_assoc] - _ ≤ A.card * (A.card / P.parts.card + 1) := + _ = #P.parts * ↑(#A / #P.parts) * ↑(#A / #P.parts + 1) := by rw [mul_assoc] + _ ≤ #A * (#A / #P.parts + 1) := mul_le_mul (mod_cast Nat.mul_div_le _ _) ?_ (by positivity) (by positivity) _ = _ := by rw [← div_add_same (mod_cast h.card_pos.ne'), mul_div_assoc] · simpa using Nat.cast_div_le - suffices (U.card - 1) * U.card ≤ A.card / P.parts.card * (A.card / P.parts.card + 1) by + suffices (#U - 1) * #U ≤ #A / #P.parts * (#A / #P.parts + 1) by rwa [Nat.mul_sub_right_distrib, one_mul, ← offDiag_card] at this have := hP.card_part_le_average_add_one hU refine Nat.mul_le_mul ((Nat.sub_le_sub_right this 1).trans ?_) this simp only [Nat.add_succ_sub_one, add_zero, card_univ, le_rfl] lemma IsEquipartition.card_biUnion_offDiag_le (hε : 0 < ε) (hP : P.IsEquipartition) - (hP' : 4 / ε ≤ P.parts.card) : (P.parts.biUnion offDiag).card ≤ ε / 2 * A.card ^ 2 := by + (hP' : 4 / ε ≤ #P.parts) : #(P.parts.biUnion offDiag) ≤ ε / 2 * #A ^ 2 := by obtain rfl | hA : A = ⊥ ∨ _ := A.eq_empty_or_nonempty · simp [Subsingleton.elim P ⊥] apply hP.card_biUnion_offDiag_le'.trans rw [div_le_iff₀ (Nat.cast_pos.2 (P.parts_nonempty hA.ne_empty).card_pos)] - have : (A.card : 𝕜) + P.parts.card ≤ 2 * A.card := by + have : (#A : 𝕜) + #P.parts ≤ 2 * #A := by rw [two_mul]; exact add_le_add_left (Nat.cast_le.2 P.card_parts_le_card) _ refine (mul_le_mul_of_nonneg_left this <| by positivity).trans ?_ - suffices 1 ≤ ε/4 * P.parts.card by + suffices 1 ≤ ε/4 * #P.parts by rw [mul_left_comm, ← sq] - convert mul_le_mul_of_nonneg_left this (mul_nonneg zero_le_two <| sq_nonneg (A.card : 𝕜)) + convert mul_le_mul_of_nonneg_left this (mul_nonneg zero_le_two <| sq_nonneg (#A : 𝕜)) using 1 <;> ring rwa [← div_le_iff₀', one_div_div] positivity lemma IsEquipartition.sum_nonUniforms_lt' (hA : A.Nonempty) (hε : 0 < ε) (hP : P.IsEquipartition) (hG : P.IsUniform G ε) : - ∑ i ∈ P.nonUniforms G ε, (i.1.card * i.2.card : 𝕜) < ε * (A.card + P.parts.card) ^ 2 := by + ∑ i ∈ P.nonUniforms G ε, (#i.1 * #i.2 : 𝕜) < ε * (#A + #P.parts) ^ 2 := by calc - _ ≤ (P.nonUniforms G ε).card • (↑(A.card / P.parts.card + 1) : 𝕜) ^ 2 := + _ ≤ #(P.nonUniforms G ε) • (↑(#A / #P.parts + 1) : 𝕜) ^ 2 := sum_le_card_nsmul _ _ _ ?_ _ = _ := nsmul_eq_mul _ _ _ ≤ _ := mul_le_mul_of_nonneg_right hG <| by positivity @@ -361,12 +360,12 @@ lemma IsEquipartition.sum_nonUniforms_lt' (hA : A.Nonempty) (hε : 0 < ε) (hP : lemma IsEquipartition.sum_nonUniforms_lt (hA : A.Nonempty) (hε : 0 < ε) (hP : P.IsEquipartition) (hG : P.IsUniform G ε) : - ((P.nonUniforms G ε).biUnion fun (U, V) ↦ U ×ˢ V).card < 4 * ε * A.card ^ 2 := by + #((P.nonUniforms G ε).biUnion fun (U, V) ↦ U ×ˢ V) < 4 * ε * #A ^ 2 := by calc - _ ≤ ∑ i ∈ P.nonUniforms G ε, (i.1.card * i.2.card : 𝕜) := by + _ ≤ ∑ i ∈ P.nonUniforms G ε, (#i.1 * #i.2 : 𝕜) := by norm_cast; simp_rw [← card_product]; exact card_biUnion_le _ < _ := hP.sum_nonUniforms_lt' hA hε hG - _ ≤ ε * (A.card + A.card) ^ 2 := by gcongr; exact P.card_parts_le_card + _ ≤ ε * (#A + #A) ^ 2 := by gcongr; exact P.card_parts_le_card _ = _ := by ring end Finpartition diff --git a/Mathlib/Combinatorics/SimpleGraph/StronglyRegular.lean b/Mathlib/Combinatorics/SimpleGraph/StronglyRegular.lean index f4078873ab459..8b0a0607d5484 100644 --- a/Mathlib/Combinatorics/SimpleGraph/StronglyRegular.lean +++ b/Mathlib/Combinatorics/SimpleGraph/StronglyRegular.lean @@ -77,7 +77,7 @@ theorem IsSRGWith.top : of_not_adj := fun v w h h' => False.elim (h' ((top_adj v w).2 h)) theorem IsSRGWith.card_neighborFinset_union_eq {v w : V} (h : G.IsSRGWith n k ℓ μ) : - (G.neighborFinset v ∪ G.neighborFinset w).card = + #(G.neighborFinset v ∪ G.neighborFinset w) = 2 * k - Fintype.card (G.commonNeighbors v w) := by apply Nat.add_right_cancel (m := Fintype.card (G.commonNeighbors v w)) rw [Nat.sub_add_cancel, ← Set.toFinset_card] @@ -94,12 +94,12 @@ adjacent to either `v` or `w` when `¬G.Adj v w`. So it's the cardinality of `G.neighborSet v ∪ G.neighborSet w`. -/ theorem IsSRGWith.card_neighborFinset_union_of_not_adj {v w : V} (h : G.IsSRGWith n k ℓ μ) (hne : v ≠ w) (ha : ¬G.Adj v w) : - (G.neighborFinset v ∪ G.neighborFinset w).card = 2 * k - μ := by + #(G.neighborFinset v ∪ G.neighborFinset w) = 2 * k - μ := by rw [← h.of_not_adj hne ha] apply h.card_neighborFinset_union_eq theorem IsSRGWith.card_neighborFinset_union_of_adj {v w : V} (h : G.IsSRGWith n k ℓ μ) - (ha : G.Adj v w) : (G.neighborFinset v ∪ G.neighborFinset w).card = 2 * k - ℓ := by + (ha : G.Adj v w) : #(G.neighborFinset v ∪ G.neighborFinset w) = 2 * k - ℓ := by rw [← h.of_adj v w ha] apply h.card_neighborFinset_union_eq @@ -172,11 +172,7 @@ theorem IsSRGWith.param_eq · simp [h.compl.regular v] · intro w hw rw [mem_neighborFinset] at hw - simp_rw [bipartiteAbove] - -- This used to be part of the enclosing `simp_rw` chain, - -- but after leanprover/lean4#3124 it caused a maximum recursion depth error. - change Finset.card (filter (fun a => Adj G w a) _) = _ - simp_rw [← mem_neighborFinset, filter_mem_eq_inter] + simp_rw [bipartiteAbove, ← mem_neighborFinset, filter_mem_eq_inter] have s : {v} ⊆ G.neighborFinset w \ G.neighborFinset v := by rw [singleton_subset_iff, mem_sdiff, mem_neighborFinset] exact ⟨hw.symm, G.not_mem_neighborFinset_self v⟩ diff --git a/Mathlib/Combinatorics/SimpleGraph/Triangle/Basic.lean b/Mathlib/Combinatorics/SimpleGraph/Triangle/Basic.lean index 96cd101472ef9..a5d36c33b508b 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Triangle/Basic.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Triangle/Basic.lean @@ -133,19 +133,19 @@ alias ⟨EdgeDisjointTriangles.mem_sym2_subsingleton, _⟩ := variable [DecidableEq α] [Fintype α] [DecidableRel G.Adj] instance EdgeDisjointTriangles.instDecidable : Decidable G.EdgeDisjointTriangles := - decidable_of_iff ((G.cliqueFinset 3 : Set (Finset α)).Pairwise fun x y ↦ ((x ∩ y).card ≤ 1)) <| by + decidable_of_iff ((G.cliqueFinset 3 : Set (Finset α)).Pairwise fun x y ↦ (#(x ∩ y) ≤ 1)) <| by simp only [coe_cliqueFinset, EdgeDisjointTriangles, Finset.card_le_one, ← coe_inter]; rfl instance LocallyLinear.instDecidable : Decidable G.LocallyLinear := inferInstanceAs (Decidable (_ ∧ _)) lemma EdgeDisjointTriangles.card_edgeFinset_le (hG : G.EdgeDisjointTriangles) : - 3 * (G.cliqueFinset 3).card ≤ G.edgeFinset.card := by - rw [mul_comm, ← mul_one G.edgeFinset.card] + 3 * #(G.cliqueFinset 3) ≤ #G.edgeFinset := by + rw [mul_comm, ← mul_one #G.edgeFinset] refine card_mul_le_card_mul (fun s e ↦ e ∈ s.sym2) ?_ (fun e he ↦ ?_) · simp only [is3Clique_iff, mem_cliqueFinset_iff, mem_sym2_iff, forall_exists_index, and_imp] rintro _ a b c hab hac hbc rfl - have : Finset.card ({s(a, b), s(a, c), s(b, c)} : Finset (Sym2 α)) = 3 := by + have : #{s(a, b), s(a, c), s(b, c)} = 3 := by refine card_eq_three.2 ⟨_, _, _, ?_, ?_, ?_, rfl⟩ <;> simp [hab.ne, hac.ne, hbc.ne] rw [← this] refine card_mono ?_ @@ -155,9 +155,9 @@ lemma EdgeDisjointTriangles.card_edgeFinset_le (hG : G.EdgeDisjointTriangles) : using hG.mem_sym2_subsingleton (G.not_isDiag_of_mem_edgeSet <| mem_edgeFinset.1 he) lemma LocallyLinear.card_edgeFinset (hG : G.LocallyLinear) : - G.edgeFinset.card = 3 * (G.cliqueFinset 3).card := by + #G.edgeFinset = 3 * #(G.cliqueFinset 3) := by refine hG.edgeDisjointTriangles.card_edgeFinset_le.antisymm' ?_ - rw [← mul_comm, ← mul_one (Finset.card _)] + rw [← mul_comm, ← mul_one #_] refine card_mul_le_card_mul (fun e s ↦ e ∈ s.sym2) ?_ ?_ · simpa [Sym2.forall, Nat.one_le_iff_ne_zero, -Finset.card_eq_zero, Finset.card_ne_zero, Finset.Nonempty] @@ -165,7 +165,7 @@ lemma LocallyLinear.card_edgeFinset (hG : G.LocallyLinear) : simp only [mem_cliqueFinset_iff, is3Clique_iff, forall_exists_index, and_imp] rintro _ a b c hab hac hbc rfl calc - _ ≤ ({s(a, b), s(a, c), s(b, c)} : Finset _).card := card_le_card ?_ + _ ≤ #{s(a, b), s(a, c), s(b, c)} := card_le_card ?_ _ ≤ 3 := (card_insert_le _ _).trans (succ_le_succ <| (card_insert_le _ _).trans_eq <| by rw [card_singleton]) simp only [subset_iff, Sym2.forall, mem_sym2_iff, le_eq_subset, mem_bipartiteBelow, mem_insert, @@ -186,7 +186,7 @@ variable {G ε} theorem farFromTriangleFree_iff : G.FarFromTriangleFree ε ↔ ∀ ⦃H : SimpleGraph α⦄, [DecidableRel H.Adj] → H ≤ G → H.CliqueFree 3 → - ε * (card α ^ 2 : ℕ) ≤ G.edgeFinset.card - H.edgeFinset.card := deleteFar_iff + ε * (card α ^ 2 : ℕ) ≤ #G.edgeFinset - #H.edgeFinset := deleteFar_iff alias ⟨farFromTriangleFree.le_card_sub_card, _⟩ := farFromTriangleFree_iff @@ -198,7 +198,7 @@ section DecidableEq variable [DecidableEq α] theorem FarFromTriangleFree.cliqueFinset_nonempty' (hH : H ≤ G) (hG : G.FarFromTriangleFree ε) - (hcard : G.edgeFinset.card - H.edgeFinset.card < ε * (card α ^ 2 : ℕ)) : + (hcard : #G.edgeFinset - #H.edgeFinset < ε * (card α ^ 2 : ℕ)) : (H.cliqueFinset 3).Nonempty := nonempty_of_ne_empty <| cliqueFinset_eq_empty_iff.not.2 fun hH' => (hG.le_card_sub_card hH hH').not_lt hcard @@ -206,7 +206,7 @@ theorem FarFromTriangleFree.cliqueFinset_nonempty' (hH : H ≤ G) (hG : G.FarFro private lemma farFromTriangleFree_of_disjoint_triangles_aux {tris : Finset (Finset α)} (htris : tris ⊆ G.cliqueFinset 3) (pd : (tris : Set (Finset α)).Pairwise fun x y ↦ (x ∩ y : Set α).Subsingleton) (hHG : H ≤ G) - (hH : H.CliqueFree 3) : tris.card ≤ G.edgeFinset.card - H.edgeFinset.card := by + (hH : H.CliqueFree 3) : #tris ≤ #G.edgeFinset - #H.edgeFinset := by rw [← card_sdiff (edgeFinset_mono hHG), ← card_attach] by_contra! hG have ⦃t⦄ (ht : t ∈ tris) : @@ -234,7 +234,7 @@ triangle-free. -/ lemma farFromTriangleFree_of_disjoint_triangles (tris : Finset (Finset α)) (htris : tris ⊆ G.cliqueFinset 3) (pd : (tris : Set (Finset α)).Pairwise fun x y ↦ (x ∩ y : Set α).Subsingleton) - (tris_big : ε * (card α ^ 2 : ℕ) ≤ tris.card) : + (tris_big : ε * (card α ^ 2 : ℕ) ≤ #tris) : G.FarFromTriangleFree ε := by rw [farFromTriangleFree_iff] intros H _ hG hH @@ -243,7 +243,7 @@ lemma farFromTriangleFree_of_disjoint_triangles (tris : Finset (Finset α)) (Nat.cast_le.2 <| farFromTriangleFree_of_disjoint_triangles_aux htris pd hG hH) protected lemma EdgeDisjointTriangles.farFromTriangleFree (hG : G.EdgeDisjointTriangles) - (tris_big : ε * (card α ^ 2 : ℕ) ≤ (G.cliqueFinset 3).card) : + (tris_big : ε * (card α ^ 2 : ℕ) ≤ #(G.cliqueFinset 3)) : G.FarFromTriangleFree ε := farFromTriangleFree_of_disjoint_triangles _ Subset.rfl (by simpa using hG) tris_big @@ -258,16 +258,16 @@ lemma FarFromTriangleFree.lt_half (hG : G.FarFromTriangleFree ε) : ε < 2⁻¹ have hε₀ : 0 < ε := hε.trans_lt' (by norm_num) rw [inv_le_iff_one_le_mul₀ (zero_lt_two' 𝕜)] at hε calc - _ ≤ (G.edgeFinset.card : 𝕜) := by + _ ≤ (#G.edgeFinset : 𝕜) := by simpa using hG.le_card_sub_card bot_le (cliqueFree_bot (le_succ _)) - _ ≤ ε * 2 * (edgeFinset G).card := le_mul_of_one_le_left (by positivity) (by assumption) + _ ≤ ε * 2 * #G.edgeFinset := le_mul_of_one_le_left (by positivity) (by assumption) _ < ε * card α ^ 2 := ?_ rw [mul_assoc, mul_lt_mul_left hε₀] norm_cast calc _ ≤ 2 * (⊤ : SimpleGraph α).edgeFinset.card := by gcongr; exact le_top _ < card α ^ 2 := ?_ - rw [edgeFinset_top, filter_not, card_sdiff (subset_univ _), card_univ, Sym2.card,] + rw [edgeFinset_top, filter_not, card_sdiff (subset_univ _), card_univ, Sym2.card] simp_rw [choose_two_right, Nat.add_sub_cancel, Nat.mul_comm _ (card α), funext (propext <| Sym2.isDiag_iff_mem_range_diag ·), univ_filter_mem_range, mul_tsub, Nat.mul_div_cancel' (card α).even_mul_succ_self.two_dvd] diff --git a/Mathlib/Combinatorics/SimpleGraph/Triangle/Counting.lean b/Mathlib/Combinatorics/SimpleGraph/Triangle/Counting.lean index f1687fb68c6d4..bff996c3fa38c 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Triangle/Counting.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Triangle/Counting.lean @@ -29,15 +29,15 @@ namespace SimpleGraph /-- The vertices of `s` whose density in `t` is `ε` less than expected. -/ private noncomputable def badVertices (ε : ℝ) (s t : Finset α) : Finset α := - s.filter fun x ↦ (t.filter <| G.Adj x).card < (G.edgeDensity s t - ε) * t.card + {x ∈ s | #{y ∈ t | G.Adj x y} < (G.edgeDensity s t - ε) * #t} private lemma card_interedges_badVertices_le : - (Rel.interedges G.Adj (badVertices G ε s t) t).card ≤ - (badVertices G ε s t).card * t.card * (G.edgeDensity s t - ε) := by + #(Rel.interedges G.Adj (badVertices G ε s t) t) ≤ + #(badVertices G ε s t) * #t * (G.edgeDensity s t - ε) := by classical refine (Nat.cast_le.2 <| (card_le_card <| subset_of_eq (Rel.interedges_eq_biUnion _)).trans card_biUnion_le).trans ?_ - simp_rw [Nat.cast_sum, card_map, ← nsmul_eq_mul, smul_mul_assoc, mul_comm (t.card : ℝ)] + simp_rw [Nat.cast_sum, card_map, ← nsmul_eq_mul, smul_mul_assoc, mul_comm (#t : ℝ)] exact sum_le_card_nsmul _ _ _ fun x hx ↦ (mem_filter.1 hx).2.le private lemma edgeDensity_badVertices_le (hε : 0 ≤ ε) (dst : 2 * ε ≤ G.edgeDensity s t) : @@ -49,7 +49,7 @@ private lemma edgeDensity_badVertices_le (hε : 0 ≤ ε) (dst : 2 * ε ≤ G.ed exact G.card_interedges_badVertices_le private lemma card_badVertices_le (dst : 2 * ε ≤ G.edgeDensity s t) (hst : G.IsUniform ε s t) : - (badVertices G ε s t).card ≤ s.card * ε := by + #(badVertices G ε s t) ≤ #s * ε := by have hε : ε ≤ 1 := (le_mul_of_one_le_of_le_of_nonneg (by norm_num) le_rfl hst.pos.le).trans (dst.trans <| by exact_mod_cast edgeDensity_le_one _ _ _) by_contra! h @@ -61,7 +61,7 @@ private lemma card_badVertices_le (dst : 2 * ε ≤ G.edgeDensity s t) (hst : G. /-- A subset of the triangles constructed in a weird way to make them easy to count. -/ private lemma triangle_split_helper [DecidableEq α] : (s \ (badVertices G ε s t ∪ badVertices G ε s u)).biUnion - (fun x ↦ (G.interedges (t.filter <| G.Adj x) (u.filter <| G.Adj x)).image (x, ·)) ⊆ + (fun x ↦ (G.interedges {y ∈ t | G.Adj x y} {y ∈ u | G.Adj x y}).image (x, ·)) ⊆ (s ×ˢ t ×ˢ u).filter (fun (x, y, z) ↦ G.Adj x y ∧ G.Adj x z ∧ G.Adj y z) := by rintro ⟨x, y, z⟩ simp only [mem_filter, mem_product, mem_biUnion, mem_sdiff, exists_prop, mem_union, @@ -72,19 +72,19 @@ private lemma triangle_split_helper [DecidableEq α] : private lemma good_vertices_triangle_card [DecidableEq α] (dst : 2 * ε ≤ G.edgeDensity s t) (dsu : 2 * ε ≤ G.edgeDensity s u) (dtu : 2 * ε ≤ G.edgeDensity t u) (utu : G.IsUniform ε t u) (x : α) (hx : x ∈ s \ (badVertices G ε s t ∪ badVertices G ε s u)) : - ε ^ 3 * t.card * u.card ≤ (((t.filter (G.Adj x) ×ˢ u.filter (G.Adj x)).filter - fun (y, z) ↦ G.Adj y z).image (x, ·)).card := by + ε ^ 3 * #t * #u ≤ #((({y ∈ t | G.Adj x y} ×ˢ {y ∈ u | G.Adj x y}).filter + fun (y, z) ↦ G.Adj y z).image (x, ·)) := by simp only [mem_sdiff, badVertices, mem_union, not_or, mem_filter, not_and_or, not_lt] at hx rw [← or_and_left, and_or_left] at hx simp only [false_or, and_not_self, mul_comm (_ - _)] at hx obtain ⟨-, hxY, hsu⟩ := hx - have hY : t.card * ε ≤ (filter (G.Adj x) t).card := + have hY : #t * ε ≤ #{y ∈ t | G.Adj x y} := (mul_le_mul_of_nonneg_left (by linarith) (Nat.cast_nonneg _)).trans hxY - have hZ : u.card * ε ≤ (filter (G.Adj x) u).card := + have hZ : #u * ε ≤ #{y ∈ u | G.Adj x y} := (mul_le_mul_of_nonneg_left (by linarith) (Nat.cast_nonneg _)).trans hsu rw [card_image_of_injective _ (Prod.mk.inj_left _)] have := utu (filter_subset (G.Adj x) _) (filter_subset (G.Adj x) _) hY hZ - have : ε ≤ G.edgeDensity (filter (G.Adj x) t) (filter (G.Adj x) u) := by + have : ε ≤ G.edgeDensity {y ∈ t | G.Adj x y} {y ∈ u | G.Adj x y} := by rw [abs_sub_lt_iff] at this; linarith rw [edgeDensity_def] at this push_cast at this @@ -101,11 +101,11 @@ lemma triangle_counting' (dst : 2 * ε ≤ G.edgeDensity s t) (hst : G.IsUniform ε s t) (dsu : 2 * ε ≤ G.edgeDensity s u) (usu : G.IsUniform ε s u) (dtu : 2 * ε ≤ G.edgeDensity t u) (utu : G.IsUniform ε t u) : - (1 - 2 * ε) * ε ^ 3 * s.card * t.card * u.card ≤ - ((s ×ˢ t ×ˢ u).filter fun (a, b, c) ↦ G.Adj a b ∧ G.Adj a c ∧ G.Adj b c).card := by + (1 - 2 * ε) * ε ^ 3 * #s * #t * #u ≤ + #((s ×ˢ t ×ˢ u).filter fun (a, b, c) ↦ G.Adj a b ∧ G.Adj a c ∧ G.Adj b c) := by classical - have h₁ : (badVertices G ε s t).card ≤ s.card * ε := G.card_badVertices_le dst hst - have h₂ : (badVertices G ε s u).card ≤ s.card * ε := G.card_badVertices_le dsu usu + have h₁ : #(badVertices G ε s t) ≤ #s * ε := G.card_badVertices_le dst hst + have h₂ : #(badVertices G ε s u) ≤ #s * ε := G.card_badVertices_le dsu usu let X' := s \ (badVertices G ε s t ∪ badVertices G ε s u) have : X'.biUnion _ ⊆ (s ×ˢ t ×ˢ u).filter fun (a, b, c) ↦ G.Adj a b ∧ G.Adj a c ∧ G.Adj b c := triangle_split_helper _ @@ -114,7 +114,7 @@ lemma triangle_counting' · apply le_trans _ (card_nsmul_le_sum X' _ _ <| G.good_vertices_triangle_card dst dsu dtu utu) rw [nsmul_eq_mul] have := hst.pos.le - suffices hX' : (1 - 2 * ε) * s.card ≤ X'.card by + suffices hX' : (1 - 2 * ε) * #s ≤ #X' by exact Eq.trans_le (by ring) (mul_le_mul_of_nonneg_right hX' <| by positivity) have i : badVertices G ε s t ∪ badVertices G ε s u ⊆ s := union_subset (filter_subset _ _) (filter_subset _ _) @@ -155,7 +155,7 @@ lemma triangle_counting (dst : 2 * ε ≤ G.edgeDensity s t) (ust : G.IsUniform ε s t) (hst : Disjoint s t) (dsu : 2 * ε ≤ G.edgeDensity s u) (usu : G.IsUniform ε s u) (hsu : Disjoint s u) (dtu : 2 * ε ≤ G.edgeDensity t u) (utu : G.IsUniform ε t u) (htu : Disjoint t u) : - (1 - 2 * ε) * ε ^ 3 * s.card * t.card * u.card ≤ (G.cliqueFinset 3).card := by + (1 - 2 * ε) * ε ^ 3 * #s * #t * #u ≤ #(G.cliqueFinset 3) := by apply (G.triangle_counting' dst ust dsu usu dtu utu).trans _ rw [Nat.cast_le] refine card_le_card_of_injOn (fun (x, y, z) ↦ {x, y, z}) ?_ ?_ diff --git a/Mathlib/Combinatorics/SimpleGraph/Triangle/Removal.lean b/Mathlib/Combinatorics/SimpleGraph/Triangle/Removal.lean index c5d8d7059920e..b212ba252ae67 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Triangle/Removal.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Triangle/Removal.lean @@ -53,21 +53,21 @@ private lemma aux {n k : ℕ} (hk : 0 < k) (hn : k ≤ n) : n < 2 * k * (n / k) apply (mod_lt n hk).trans_le simpa using Nat.mul_le_mul_left k ((Nat.one_le_div_iff hk).2 hn) -private lemma card_bound (hP₁ : P.IsEquipartition) (hP₃ : P.parts.card ≤ bound (ε / 8) ⌈4/ε⌉₊) - (hX : s ∈ P.parts) : card α / (2 * bound (ε / 8) ⌈4 / ε⌉₊ : ℝ) ≤ s.card := by +private lemma card_bound (hP₁ : P.IsEquipartition) (hP₃ : #P.parts ≤ bound (ε / 8) ⌈4/ε⌉₊) + (hX : s ∈ P.parts) : card α / (2 * bound (ε / 8) ⌈4 / ε⌉₊ : ℝ) ≤ #s := by cases isEmpty_or_nonempty α · simp [Fintype.card_eq_zero] have := Finset.Nonempty.card_pos ⟨_, hX⟩ calc - _ ≤ card α / (2 * P.parts.card : ℝ) := by gcongr - _ ≤ ↑(card α / P.parts.card) := + _ ≤ card α / (2 * #P.parts : ℝ) := by gcongr + _ ≤ ↑(card α / #P.parts) := (div_le_iff₀' (by positivity)).2 <| mod_cast (aux ‹_› P.card_parts_le_card).le - _ ≤ (s.card : ℝ) := mod_cast hP₁.average_le_card_part hX + _ ≤ (#s : ℝ) := mod_cast hP₁.average_le_card_part hX private lemma triangle_removal_aux (hε : 0 < ε) (hP₁ : P.IsEquipartition) - (hP₃ : P.parts.card ≤ bound (ε / 8) ⌈4/ε⌉₊) + (hP₃ : #P.parts ≤ bound (ε / 8) ⌈4/ε⌉₊) (ht : t ∈ (G.regularityReduced P (ε/8) (ε/4)).cliqueFinset 3) : - triangleRemovalBound ε * card α ^ 3 ≤ (G.cliqueFinset 3).card := by + triangleRemovalBound ε * card α ^ 3 ≤ #(G.cliqueFinset 3) := by rw [mem_cliqueFinset_iff, is3Clique_iff] at ht obtain ⟨x, y, z, ⟨-, s, hX, Y, hY, xX, yY, nXY, uXY, dXY⟩, ⟨-, X', hX', Z, hZ, xX', zZ, nXZ, uXZ, dXZ⟩, @@ -87,28 +87,28 @@ private lemma triangle_removal_aux (hε : 0 < ε) (hP₁ : P.IsEquipartition) _ = (1 - 2 * (ε / 8)) * (ε / 8) ^ 3 * (card α / (2 * bound (ε / 8) ⌈4 / ε⌉₊)) * (card α / (2 * bound (ε / 8) ⌈4 / ε⌉₊)) * (card α / (2 * bound (ε / 8) ⌈4 / ε⌉₊)) := by ring - _ ≤ (1 - 2 * (ε / 8)) * (ε / 8) ^ 3 * s.card * Y.card * Z.card := by + _ ≤ (1 - 2 * (ε / 8)) * (ε / 8) ^ 3 * #s * #Y * #Z := by gcongr <;> exact card_bound hP₁ hP₃ ‹_› _ ≤ _ := triangle_counting G (by rwa [that]) uXY dXY (by rwa [that]) uXZ dXZ (by rwa [that]) uYZ dYZ lemma regularityReduced_edges_card_aux [Nonempty α] (hε : 0 < ε) (hP : P.IsEquipartition) - (hPε : P.IsUniform G (ε/8)) (hP' : 4 / ε ≤ P.parts.card) : - 2 * (G.edgeFinset.card - (G.regularityReduced P (ε/8) (ε/4)).edgeFinset.card : ℝ) + (hPε : P.IsUniform G (ε/8)) (hP' : 4 / ε ≤ #P.parts) : + 2 * (#G.edgeFinset - #(G.regularityReduced P (ε/8) (ε/4)).edgeFinset : ℝ) < 2 * ε * (card α ^ 2 : ℕ) := by let A := (P.nonUniforms G (ε/8)).biUnion fun (U, V) ↦ U ×ˢ V let B := P.parts.biUnion offDiag let C := (P.sparsePairs G (ε/4)).biUnion fun (U, V) ↦ G.interedges U V calc - _ = ↑((univ ×ˢ univ).filter fun (x, y) ↦ - G.Adj x y ∧ ¬(G.regularityReduced P (ε / 8) (ε /4)).Adj x y).card := by + _ = (#((univ ×ˢ univ).filter fun (x, y) ↦ + G.Adj x y ∧ ¬(G.regularityReduced P (ε / 8) (ε /4)).Adj x y) : ℝ) := by rw [univ_product_univ, mul_sub, filter_and_not, cast_card_sdiff] · norm_cast rw [two_mul_card_edgeFinset, two_mul_card_edgeFinset] · exact monotone_filter_right _ fun xy hxy ↦ regularityReduced_le hxy - _ ≤ ((A ∪ B ∪ C).card : ℝ) := by gcongr; exact unreduced_edges_subset - _ ≤ ((A ∪ B).card + C.card : ℝ) := mod_cast (card_union_le _ _) - _ ≤ (A.card + B.card + C.card : ℝ) := by gcongr; exact mod_cast card_union_le _ _ + _ ≤ #(A ∪ B ∪ C) := by gcongr; exact unreduced_edges_subset + _ ≤ #(A ∪ B) + #C := mod_cast (card_union_le _ _) + _ ≤ #A + #B + #C := by gcongr; exact mod_cast card_union_le _ _ _ < 4 * (ε / 8) * card α ^ 2 + _ + _ := by gcongr; exact hP.sum_nonUniforms_lt univ_nonempty (by positivity) hPε _ ≤ _ + ε / 2 * card α ^ 2 + 4 * (ε / 4) * card α ^ 2 := by @@ -121,7 +121,7 @@ lemma regularityReduced_edges_card_aux [Nonempty α] (hε : 0 < ε) (hP : P.IsEq order of `(card α)^2`), then there were many triangles to start with (on the order of `(card α)^3`). -/ lemma FarFromTriangleFree.le_card_cliqueFinset (hG : G.FarFromTriangleFree ε) : - triangleRemovalBound ε * card α ^ 3 ≤ (G.cliqueFinset 3).card := by + triangleRemovalBound ε * card α ^ 3 ≤ #(G.cliqueFinset 3) := by cases isEmpty_or_nonempty α · simp [Fintype.card_eq_zero] obtain hε | hε := le_or_lt ε 0 @@ -135,7 +135,7 @@ lemma FarFromTriangleFree.le_card_cliqueFinset (hG : G.FarFromTriangleFree ε) : _ ≤ (1 : ℝ) := (triangleRemovalBound_mul_cube_lt hε).le _ ≤ _ := by simpa [one_le_iff_ne_zero] using (hG.cliqueFinset_nonempty hε).card_pos.ne' obtain ⟨P, hP₁, hP₂, hP₃, hP₄⟩ := szemeredi_regularity G (by positivity : 0 < ε / 8) hl' - have : 4/ε ≤ P.parts.card := hl.trans (cast_le.2 hP₂) + have : 4/ε ≤ #P.parts := hl.trans (cast_le.2 hP₂) have k := regularityReduced_edges_card_aux hε hP₁ hP₄ this rw [mul_assoc] at k replace k := lt_of_mul_lt_mul_left k zero_le_two @@ -144,9 +144,9 @@ lemma FarFromTriangleFree.le_card_cliqueFinset (hG : G.FarFromTriangleFree ε) : /-- **Triangle Removal Lemma**. If there are not too many triangles (on the order of `(card α)^3`) then they can all be removed by removing a few edges (on the order of `(card α)^2`). -/ -lemma triangle_removal (hG : (G.cliqueFinset 3).card < triangleRemovalBound ε * card α ^ 3) : +lemma triangle_removal (hG : #(G.cliqueFinset 3) < triangleRemovalBound ε * card α ^ 3) : ∃ G' ≤ G, ∃ _ : DecidableRel G'.Adj, - (G.edgeFinset.card - G'.edgeFinset.card : ℝ) < ε * (card α^2 : ℕ) ∧ G'.CliqueFree 3 := by + (#G.edgeFinset - #G'.edgeFinset : ℝ) < ε * (card α^2 : ℕ) ∧ G'.CliqueFree 3 := by by_contra! h refine hG.not_le (farFromTriangleFree_iff.2 ?_).le_card_cliqueFinset intros G' _ hG hG' diff --git a/Mathlib/Combinatorics/SimpleGraph/Triangle/Tripartite.lean b/Mathlib/Combinatorics/SimpleGraph/Triangle/Tripartite.lean index 6bcfe21447840..3135218de5e3d 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Triangle/Tripartite.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Triangle/Tripartite.lean @@ -216,11 +216,11 @@ lemma cliqueFinset_eq_image [NoAccidental t] : (graph t).cliqueFinset 3 = t.imag lemma cliqueFinset_eq_map [NoAccidental t] : (graph t).cliqueFinset 3 = t.map toTriangle := by simp [cliqueFinset_eq_image, map_eq_image] -@[simp] lemma card_triangles [NoAccidental t] : ((graph t).cliqueFinset 3).card = t.card := by +@[simp] lemma card_triangles [NoAccidental t] : #((graph t).cliqueFinset 3) = #t := by rw [cliqueFinset_eq_map, card_map] lemma farFromTriangleFree [ExplicitDisjoint t] {ε : 𝕜} - (ht : ε * ((Fintype.card α + Fintype.card β + Fintype.card γ) ^ 2 : ℕ) ≤ t.card) : + (ht : ε * ((Fintype.card α + Fintype.card β + Fintype.card γ) ^ 2 : ℕ) ≤ #t) : (graph t).FarFromTriangleFree ε := farFromTriangleFree_of_disjoint_triangles (t.map toTriangle) (map_subset_iff_subset_preimage.2 fun x hx ↦ by simpa using toTriangle_is3Clique hx) diff --git a/Mathlib/Combinatorics/SimpleGraph/Turan.lean b/Mathlib/Combinatorics/SimpleGraph/Turan.lean index 08ab8e530a2c0..bb502466b40b0 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Turan.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Turan.lean @@ -49,7 +49,7 @@ variable (G) in the same vertex set has the same or fewer number of edges. -/ def IsTuranMaximal (r : ℕ) : Prop := G.CliqueFree (r + 1) ∧ ∀ (H : SimpleGraph V) [DecidableRel H.Adj], - H.CliqueFree (r + 1) → H.edgeFinset.card ≤ G.edgeFinset.card + H.CliqueFree (r + 1) → #H.edgeFinset ≤ #G.edgeFinset section Defs @@ -111,7 +111,7 @@ lemma exists_isTuranMaximal (hr : 0 < r): have cn : c.toFinset.Nonempty := ⟨⊥, by simp only [Set.toFinset_setOf, mem_filter, mem_univ, true_and, c] exact cliqueFree_bot (by omega)⟩ - obtain ⟨S, Sm, Sl⟩ := exists_max_image c.toFinset (·.edgeFinset.card) cn + obtain ⟨S, Sm, Sl⟩ := exists_max_image c.toFinset (#·.edgeFinset) cn use S, inferInstance rw [Set.mem_toFinset] at Sm refine ⟨Sm, fun I _ cf ↦ ?_⟩ @@ -195,11 +195,11 @@ lemma not_adj_iff_part_eq [DecidableEq V] : rw [fp.mem_part_iff_part_eq_part (mem_univ t) (mem_univ s), eq_comm] lemma degree_eq_card_sub_part_card [DecidableEq V] : - G.degree s = Fintype.card V - (h.finpartition.part s).card := + G.degree s = Fintype.card V - #(h.finpartition.part s) := calc - _ = (univ.filter (G.Adj s)).card := by + _ = #{t | G.Adj s t} := by simp [← card_neighborFinset_eq_degree, neighborFinset] - _ = Fintype.card V - (univ.filter (¬G.Adj s ·)).card := + _ = Fintype.card V - #{t | ¬G.Adj s t} := eq_tsub_of_add_eq (filter_card_add_filter_neg_card_eq_card _) _ = _ := by congr; ext; rw [mem_filter] @@ -224,13 +224,13 @@ theorem isEquipartition [DecidableEq V] : h.finpartition.IsEquipartition := by rw [hn] at ineq; omega rw [G.card_edgeFinset_replaceVertex_of_adj ha, degree_eq_card_sub_part_card h, small_eq, degree_eq_card_sub_part_card h, large_eq] - have : large.card ≤ Fintype.card V := by simpa using card_le_card large.subset_univ + have : #large ≤ Fintype.card V := by simpa using card_le_card large.subset_univ omega -lemma card_parts_le [DecidableEq V] : h.finpartition.parts.card ≤ r := by +lemma card_parts_le [DecidableEq V] : #h.finpartition.parts ≤ r := by by_contra! l obtain ⟨z, -, hz⟩ := h.finpartition.exists_subset_part_bijOn - have ncf : ¬G.CliqueFree z.card := by + have ncf : ¬G.CliqueFree #z := by refine IsNClique.not_cliqueFree ⟨fun v hv w hw hn ↦ ?_, rfl⟩ contrapose! hn exact hz.injOn hv hw (by rwa [← h.not_adj_iff_part_eq]) @@ -240,7 +240,7 @@ lemma card_parts_le [DecidableEq V] : h.finpartition.parts.card ≤ r := by /-- There are `min n r` parts in a graph on `n` vertices satisfying `G.IsTuranMaximal r`. `min` handles the `n < r` case, when `G` is complete but still `r + 1`-cliquefree for having insufficiently many vertices. -/ -theorem card_parts [DecidableEq V] : h.finpartition.parts.card = min (Fintype.card V) r := by +theorem card_parts [DecidableEq V] : #h.finpartition.parts = min (Fintype.card V) r := by set fp := h.finpartition apply le_antisymm (le_min fp.card_parts_le_card h.card_parts_le) by_contra! l @@ -255,7 +255,7 @@ theorem card_parts [DecidableEq V] : h.finpartition.parts.card = min (Fintype.ca intro z zc; push_neg; simp_rw [h.not_adj_iff_part_eq] exact exists_ne_map_eq_of_card_lt_of_maps_to (zc.symm ▸ l.2) fun a _ ↦ fp.part_mem (mem_univ a) use G ⊔ edge x y, inferInstance, cf.sup_edge x y - convert Nat.lt.base G.edgeFinset.card + convert Nat.lt.base #G.edgeFinset convert G.card_edgeFinset_sup_edge _ hn rwa [h.not_adj_iff_part_eq] From 9a2abc12f24fe15a8a808fd0fcc39520b1b2e0a7 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Tue, 22 Oct 2024 14:15:51 +0000 Subject: [PATCH 274/505] feat(SetTheory/Cardinal/Basic): `lift (aleph' o) = aleph' (lift o)` (#16306) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit plus some lemmas on `ℵ₁` and `lift`. --- Mathlib/SetTheory/Cardinal/Aleph.lean | 34 +++++++++++++++++++++++++++ Mathlib/SetTheory/Cardinal/Basic.lean | 16 +++++++++---- 2 files changed, 46 insertions(+), 4 deletions(-) diff --git a/Mathlib/SetTheory/Cardinal/Aleph.lean b/Mathlib/SetTheory/Cardinal/Aleph.lean index 090dd44cb90c6..d41e56c312df0 100644 --- a/Mathlib/SetTheory/Cardinal/Aleph.lean +++ b/Mathlib/SetTheory/Cardinal/Aleph.lean @@ -233,6 +233,12 @@ theorem aleph'_nat : ∀ n : ℕ, aleph' n = n | 0 => aleph'_zero | n + 1 => show aleph' (succ n) = n.succ by rw [aleph'_succ, aleph'_nat n, nat_succ] +set_option linter.docPrime false in +@[simp] +theorem lift_aleph' (o : Ordinal.{u}) : lift.{v} (aleph' o) = aleph' (Ordinal.lift.{v} o) := + ((InitialSeg.ofIso aleph'.toRelIsoLT).trans liftInitialSeg).eq + (Ordinal.liftInitialSeg.trans (InitialSeg.ofIso aleph'.toRelIsoLT)) o + theorem aleph'_le_of_limit {o : Ordinal} (l : o.IsLimit) {c} : aleph' o ≤ c ↔ ∀ o' < o, aleph' o' ≤ c := ⟨fun h o' h' => (aleph'_le.2 <| h'.le).trans h, fun h => by @@ -297,6 +303,10 @@ theorem aleph_succ (o : Ordinal) : ℵ_ (succ o) = succ (ℵ_ o) := by @[simp] theorem aleph_zero : ℵ_ 0 = ℵ₀ := by rw [aleph_eq_aleph', add_zero, aleph'_omega0] +@[simp] +theorem lift_aleph (o : Ordinal.{u}) : lift.{v} (aleph o) = aleph (Ordinal.lift.{v} o) := by + simp [aleph_eq_aleph'] + theorem aleph_limit {o : Ordinal} (ho : o.IsLimit) : ℵ_ o = ⨆ a : Iio o, ℵ_ a := by apply le_antisymm _ (ciSup_le' _) · rw [aleph_eq_aleph', aleph'_limit (ho.add _)] @@ -361,6 +371,30 @@ theorem aleph0_lt_aleph_one : ℵ₀ < ℵ₁ := by theorem countable_iff_lt_aleph_one {α : Type*} (s : Set α) : s.Countable ↔ #s < ℵ₁ := by rw [← succ_aleph0, lt_succ_iff, le_aleph0_iff_set_countable] +@[simp] +theorem aleph1_le_lift {c : Cardinal.{u}} : ℵ₁ ≤ lift.{v} c ↔ ℵ₁ ≤ c := by + simpa using lift_le (a := ℵ₁) + +@[simp] +theorem lift_le_aleph1 {c : Cardinal.{u}} : lift.{v} c ≤ ℵ₁ ↔ c ≤ ℵ₁ := by + simpa using lift_le (b := ℵ₁) + +@[simp] +theorem aleph1_lt_lift {c : Cardinal.{u}} : ℵ₁ < lift.{v} c ↔ ℵ₁ < c := by + simpa using lift_lt (a := ℵ₁) + +@[simp] +theorem lift_lt_aleph1 {c : Cardinal.{u}} : lift.{v} c < ℵ₁ ↔ c < ℵ₁ := by + simpa using lift_lt (b := ℵ₁) + +@[simp] +theorem aleph1_eq_lift {c : Cardinal.{u}} : ℵ₁ = lift.{v} c ↔ ℵ₁ = c := by + simpa using lift_inj (a := ℵ₁) + +@[simp] +theorem lift_eq_aleph1 {c : Cardinal.{u}} : lift.{v} c = ℵ₁ ↔ c = ℵ₁ := by + simpa using lift_inj (b := ℵ₁) + section deprecated set_option linter.deprecated false diff --git a/Mathlib/SetTheory/Cardinal/Basic.lean b/Mathlib/SetTheory/Cardinal/Basic.lean index fcc9e9174c5c2..90ae6898a01ca 100644 --- a/Mathlib/SetTheory/Cardinal/Basic.lean +++ b/Mathlib/SetTheory/Cardinal/Basic.lean @@ -1250,19 +1250,27 @@ theorem lift_aleph0 : lift ℵ₀ = ℵ₀ := @[simp] theorem aleph0_le_lift {c : Cardinal.{u}} : ℵ₀ ≤ lift.{v} c ↔ ℵ₀ ≤ c := by - rw [← lift_aleph0.{v, u}, lift_le] + simpa using lift_le (a := ℵ₀) @[simp] theorem lift_le_aleph0 {c : Cardinal.{u}} : lift.{v} c ≤ ℵ₀ ↔ c ≤ ℵ₀ := by - rw [← lift_aleph0.{v, u}, lift_le] + simpa using lift_le (b := ℵ₀) @[simp] theorem aleph0_lt_lift {c : Cardinal.{u}} : ℵ₀ < lift.{v} c ↔ ℵ₀ < c := by - rw [← lift_aleph0.{v, u}, lift_lt] + simpa using lift_lt (a := ℵ₀) @[simp] theorem lift_lt_aleph0 {c : Cardinal.{u}} : lift.{v} c < ℵ₀ ↔ c < ℵ₀ := by - rw [← lift_aleph0.{v, u}, lift_lt] + simpa using lift_lt (b := ℵ₀) + +@[simp] +theorem aleph0_eq_lift {c : Cardinal.{u}} : ℵ₀ = lift.{v} c ↔ ℵ₀ = c := by + simpa using lift_inj (a := ℵ₀) + +@[simp] +theorem lift_eq_aleph0 {c : Cardinal.{u}} : lift.{v} c = ℵ₀ ↔ c = ℵ₀ := by + simpa using lift_inj (b := ℵ₀) /-! ### Properties about the cast from `ℕ` -/ From b5f6dae54ae5cc33e326a3e064780517d91ea20e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Tue, 22 Oct 2024 14:15:52 +0000 Subject: [PATCH 275/505] =?UTF-8?q?chore(SetTheory/Ordinal/Arithmetic):=20?= =?UTF-8?q?`WellFounded.rank`=20=E2=86=92=20`IsWellFounded.rank`=20(#16586?= =?UTF-8?q?)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We make the `WellFounded` argument of `rank` an instance argument. This matches `Ordinal.typein`. --- Mathlib/Order/Extension/Well.lean | 45 ++++++++++++++++++--- Mathlib/SetTheory/Ordinal/Arithmetic.lean | 49 +++++++++++++++++++---- Mathlib/SetTheory/ZFC/Rank.lean | 14 +++---- 3 files changed, 88 insertions(+), 20 deletions(-) diff --git a/Mathlib/Order/Extension/Well.lean b/Mathlib/Order/Extension/Well.lean index 6c7088100dd59..f3f4d25022936 100644 --- a/Mathlib/Order/Extension/Well.lean +++ b/Mathlib/Order/Extension/Well.lean @@ -16,7 +16,7 @@ well-founded order. We can map our order into two well-orders: * the first map respects the order but isn't necessarily injective. Namely, this is the *rank* - function `WellFounded.rank : α → Ordinal`. + function `IsWellFounded.rank : α → Ordinal`. * the second map is injective but doesn't necessarily respect the order. This is an arbitrary embedding into `Cardinal` given by `embeddingToCardinal`. @@ -38,8 +38,40 @@ universe u variable {α : Type u} {r : α → α → Prop} +namespace IsWellFounded + +variable {α : Type u} (r : α → α → Prop) [IsWellFounded α r] + +/-- An arbitrary well order on `α` that extends `r`. + +The construction maps `r` into two well-orders: the first map is `IsWellFounded.rank`, which is not +necessarily injective but respects the order `r`; the other map is the identity (with an arbitrarily +chosen well-order on `α`), which is injective but doesn't respect `r`. + +By taking the lexicographic product of the two, we get both properties, so we can pull it back and +get a well-order that extend our original order `r`. Another way to view this is that we choose an +arbitrary well-order to serve as a tiebreak between two elements of same rank. +-/ +noncomputable def wellOrderExtension : LinearOrder α := + @LinearOrder.lift' α (Ordinal ×ₗ Cardinal) _ (fun a : α => (rank r a, embeddingToCardinal a)) + fun _ _ h => embeddingToCardinal.injective <| congr_arg Prod.snd h + +instance wellOrderExtension.isWellFounded_lt : IsWellFounded α (wellOrderExtension r).lt := + ⟨InvImage.wf (fun a : α => (rank r a, embeddingToCardinal a)) <| + Ordinal.lt_wf.prod_lex Cardinal.lt_wf⟩ + +instance wellOrderExtension.isWellOrder_lt : IsWellOrder α (wellOrderExtension r).lt where + +/-- Any well-founded relation can be extended to a well-ordering on that type. -/ +theorem exists_well_order_ge : ∃ s, r ≤ s ∧ IsWellOrder α s := + ⟨(wellOrderExtension r).lt, fun _ _ h => Prod.Lex.left _ _ (rank_lt_of_rel h), ⟨⟩⟩ + +end IsWellFounded + namespace WellFounded +set_option linter.deprecated false + variable (hwf : WellFounded r) /-- An arbitrary well order on `α` that extends `r`. @@ -52,16 +84,19 @@ By taking the lexicographic product of the two, we get both properties, so we ca get a well-order that extend our original order `r`. Another way to view this is that we choose an arbitrary well-order to serve as a tiebreak between two elements of same rank. -/ +@[deprecated IsWellFounded.wellOrderExtension (since := "2024-09-07")] noncomputable def wellOrderExtension : LinearOrder α := @LinearOrder.lift' α (Ordinal ×ₗ Cardinal) _ (fun a : α => (hwf.rank a, embeddingToCardinal a)) fun _ _ h => embeddingToCardinal.injective <| congr_arg Prod.snd h +@[deprecated IsWellFounded.wellOrderExtension.isWellFounded_lt (since := "2024-09-07")] instance wellOrderExtension.isWellFounded_lt : IsWellFounded α hwf.wellOrderExtension.lt := ⟨InvImage.wf (fun a : α => (hwf.rank a, embeddingToCardinal a)) <| Ordinal.lt_wf.prod_lex Cardinal.lt_wf⟩ include hwf in /-- Any well-founded relation can be extended to a well-ordering on that type. -/ +@[deprecated IsWellFounded.exists_well_order_ge (since := "2024-09-07")] theorem exists_well_order_ge : ∃ s, r ≤ s ∧ IsWellOrder α s := ⟨hwf.wellOrderExtension.lt, fun _ _ h => Prod.Lex.left _ _ (hwf.rank_lt_of_rel h), ⟨⟩⟩ @@ -76,13 +111,13 @@ instance [Inhabited α] : Inhabited (WellOrderExtension α) := ‹_› def toWellOrderExtension : α ≃ WellOrderExtension α := Equiv.refl _ -noncomputable instance [LT α] [WellFoundedLT α] : LinearOrder (WellOrderExtension α) := - (IsWellFounded.wf : @WellFounded α (· < ·)).wellOrderExtension +noncomputable instance [LT α] [h : WellFoundedLT α] : LinearOrder (WellOrderExtension α) := + h.wellOrderExtension instance WellOrderExtension.wellFoundedLT [LT α] [WellFoundedLT α] : WellFoundedLT (WellOrderExtension α) := - WellFounded.wellOrderExtension.isWellFounded_lt _ + IsWellFounded.wellOrderExtension.isWellFounded_lt (α := α) (· < ·) theorem toWellOrderExtension_strictMono [Preorder α] [WellFoundedLT α] : StrictMono (toWellOrderExtension : α → WellOrderExtension α) := fun _ _ h => - Prod.Lex.left _ _ <| WellFounded.rank_lt_of_rel _ h + Prod.Lex.left _ _ <| IsWellFounded.rank_lt_of_rel h diff --git a/Mathlib/SetTheory/Ordinal/Arithmetic.lean b/Mathlib/SetTheory/Ordinal/Arithmetic.lean index 3035879d9e2b4..074ad4247ca3a 100644 --- a/Mathlib/SetTheory/Ordinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Ordinal/Arithmetic.lean @@ -2415,10 +2415,12 @@ theorem noMaxOrder {c} (h : ℵ₀ ≤ c) : NoMaxOrder c.ord.toType := end Cardinal -variable {α : Type u} {r : α → α → Prop} {a b : α} +variable {α : Type u} {a b : α} namespace Acc +variable {r : α → α → Prop} + /-- The rank of an element `a` accessible under a relation `r` is defined inductively as the smallest ordinal greater than the ranks of all elements below it (i.e. elements `b` such that `r b a`). -/ @@ -2434,32 +2436,63 @@ theorem rank_eq (h : Acc r a) : theorem rank_lt_of_rel (hb : Acc r b) (h : r a b) : (hb.inv h).rank < hb.rank := (Order.lt_succ _).trans_le <| by rw [hb.rank_eq] - refine le_trans ?_ (Ordinal.le_iSup _ ⟨a, h⟩) - rfl + exact Ordinal.le_iSup _ (⟨a, h⟩ : {a // r a b}) end Acc +namespace IsWellFounded + +variable (r : α → α → Prop) [hwf : IsWellFounded α r] + +/-- The rank of an element `a` under a well-founded relation `r` is defined inductively as the +smallest ordinal greater than the ranks of all elements below it (i.e. elements `b` such that +`r b a`). -/ +noncomputable def rank (a : α) : Ordinal.{u} := + (hwf.apply r a).rank + +theorem rank_eq (a : α) : rank r a = ⨆ b : { b // r b a }, succ (rank r b) := + (hwf.apply r a).rank_eq + +variable {r} in +theorem rank_lt_of_rel (h : r a b) : rank r a < rank r b := + Acc.rank_lt_of_rel _ h + +end IsWellFounded + +theorem WellFoundedLT.rank_strictMono [Preorder α] [WellFoundedLT α] : + StrictMono (IsWellFounded.rank (α := α) (· < ·)) := + fun _ _ => IsWellFounded.rank_lt_of_rel + +theorem WellFoundedGT.rank_strictAnti [Preorder α] [WellFoundedGT α] : + StrictAnti (IsWellFounded.rank (α := α) (· > ·)) := + fun _ _ a => IsWellFounded.rank_lt_of_rel a + namespace WellFounded -variable (hwf : WellFounded r) +set_option linter.deprecated false + +variable {r : α → α → Prop} (hwf : WellFounded r) /-- The rank of an element `a` under a well-founded relation `r` is defined inductively as the smallest ordinal greater than the ranks of all elements below it (i.e. elements `b` such that `r b a`). -/ +@[deprecated IsWellFounded.rank (since := "2024-09-07")] noncomputable def rank (a : α) : Ordinal.{u} := (hwf.apply a).rank -theorem rank_eq : - hwf.rank a = ⨆ b : { b // r b a }, Order.succ (hwf.rank b) := by - rw [rank, Acc.rank_eq] - rfl +@[deprecated IsWellFounded.rank_eq (since := "2024-09-07")] +theorem rank_eq : hwf.rank a = ⨆ b : { b // r b a }, Order.succ (hwf.rank b) := + (hwf.apply a).rank_eq +@[deprecated IsWellFounded.rank_lt_of_rel (since := "2024-09-07")] theorem rank_lt_of_rel (h : r a b) : hwf.rank a < hwf.rank b := Acc.rank_lt_of_rel _ h +@[deprecated WellFoundedLT.rank_strictMono (since := "2024-09-07")] theorem rank_strictMono [Preorder α] [WellFoundedLT α] : StrictMono (rank <| @wellFounded_lt α _ _) := fun _ _ => rank_lt_of_rel _ +@[deprecated WellFoundedGT.rank_strictAnti (since := "2024-09-07")] theorem rank_strictAnti [Preorder α] [WellFoundedGT α] : StrictAnti (rank <| @wellFounded_gt α _ _) := fun _ _ => rank_lt_of_rel wellFounded_gt diff --git a/Mathlib/SetTheory/ZFC/Rank.lean b/Mathlib/SetTheory/ZFC/Rank.lean index a537948adaba8..fe8e38cbca3b7 100644 --- a/Mathlib/SetTheory/ZFC/Rank.lean +++ b/Mathlib/SetTheory/ZFC/Rank.lean @@ -10,7 +10,7 @@ import Mathlib.SetTheory.ZFC.Basic # Ordinal ranks of PSet and ZFSet In this file, we define the ordinal ranks of `PSet` and `ZFSet`. These ranks are the same as -`WellFounded.rank` over `∈`, but are defined in a way that the universe levels of ranks are the +`IsWellFounded.rank` over `∈`, but are defined in a way that the universe levels of ranks are the same as the indexing types. ## Definitions @@ -109,10 +109,10 @@ theorem le_succ_rank_sUnion : rank x ≤ succ (rank (⋃₀ x)) := by rw [mem_sUnion] exists z -/-- `PSet.rank` is equal to the `WellFounded.rank` over `∈`. -/ -theorem rank_eq_wfRank : lift.{u + 1, u} (rank x) = mem_wf.rank x := by +/-- `PSet.rank` is equal to the `IsWellFounded.rank` over `∈`. -/ +theorem rank_eq_wfRank : lift.{u + 1, u} (rank x) = IsWellFounded.rank (α := PSet) (· ∈ ·) x := by induction' x using mem_wf.induction with x ih - rw [mem_wf.rank_eq] + rw [IsWellFounded.rank_eq] simp_rw [← fun y : { y // y ∈ x } => ih y y.2] apply (le_of_forall_lt _).antisymm (Ordinal.iSup_le _) <;> intro h · rw [lt_lift_iff] @@ -195,10 +195,10 @@ theorem rank_range {α : Type u} {f : α → ZFSet.{max u v}} : · simpa [rank_le_iff] using lt_lsub _ · simp [rank_lt_of_mem] -/-- `ZFSet.rank` is equal to the `WellFounded.rank` over `∈`. -/ -theorem rank_eq_wfRank : lift.{u + 1, u} (rank x) = mem_wf.rank x := by +/-- `ZFSet.rank` is equal to the `IsWellFounded.rank` over `∈`. -/ +theorem rank_eq_wfRank : lift.{u + 1, u} (rank x) = IsWellFounded.rank (α := ZFSet) (· ∈ ·) x := by induction' x using inductionOn with x ih - rw [mem_wf.rank_eq] + rw [IsWellFounded.rank_eq] simp_rw [← fun y : { y // y ∈ x } => ih y y.2] apply (le_of_forall_lt _).antisymm (Ordinal.iSup_le _) <;> intro h · rw [lt_lift_iff] From f94dd660e5d3567d83def67decefae373ecb09a4 Mon Sep 17 00:00:00 2001 From: Hannah Scholz Date: Tue, 22 Oct 2024 14:15:53 +0000 Subject: [PATCH 276/505] feat: 'and' and 'exist' commute for Antitone or Monotone properties (#17947) --- Mathlib/Order/BoundedOrder.lean | 10 ++++++++++ 1 file changed, 10 insertions(+) diff --git a/Mathlib/Order/BoundedOrder.lean b/Mathlib/Order/BoundedOrder.lean index a28c66b0685a2..7c89768bd45e6 100644 --- a/Mathlib/Order/BoundedOrder.lean +++ b/Mathlib/Order/BoundedOrder.lean @@ -503,6 +503,11 @@ theorem exists_ge_and_iff_exists {P : α → Prop} {x₀ : α} (hP : Monotone P) (∃ x, x₀ ≤ x ∧ P x) ↔ ∃ x, P x := ⟨fun h => h.imp fun _ h => h.2, fun ⟨x, hx⟩ => ⟨x ⊔ x₀, le_sup_right, hP le_sup_left hx⟩⟩ +lemma exists_and_iff_of_monotone {P Q : α → Prop} (hP : Monotone P) (hQ : Monotone Q) : + ((∃ x, P x) ∧ ∃ x, Q x) ↔ (∃ x, P x ∧ Q x) := + ⟨fun ⟨⟨x, hPx⟩, ⟨y, hQx⟩⟩ ↦ ⟨x ⊔ y, ⟨hP le_sup_left hPx, hQ le_sup_right hQx⟩⟩, + fun ⟨x, hPx, hQx⟩ ↦ ⟨⟨x, hPx⟩, ⟨x, hQx⟩⟩⟩ + end SemilatticeSup section SemilatticeInf @@ -513,6 +518,11 @@ theorem exists_le_and_iff_exists {P : α → Prop} {x₀ : α} (hP : Antitone P) (∃ x, x ≤ x₀ ∧ P x) ↔ ∃ x, P x := exists_ge_and_iff_exists <| hP.dual_left +lemma exists_and_iff_of_antitone {P Q : α → Prop} (hP : Antitone P) (hQ : Antitone Q) : + ((∃ x, P x) ∧ ∃ x, Q x) ↔ (∃ x, P x ∧ Q x) := + ⟨fun ⟨⟨x, hPx⟩, ⟨y, hQx⟩⟩ ↦ ⟨x ⊓ y, ⟨hP inf_le_left hPx, hQ inf_le_right hQx⟩⟩, + fun ⟨x, hPx, hQx⟩ ↦ ⟨⟨x, hPx⟩, ⟨x, hQx⟩⟩⟩ + end SemilatticeInf end Logic From 660d68a7ddba2cb2eb0aaa67c8dc15eb9bc2a2bc Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 22 Oct 2024 14:15:55 +0000 Subject: [PATCH 277/505] =?UTF-8?q?refactor(Group/Pointwise/Finset):=20sta?= =?UTF-8?q?te=20`mul=5Fsingleton`/`singleton=5Fmul`=20using=20`=E2=80=A2`?= =?UTF-8?q?=20(#18068)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This is much more useful in practice than the `image` spelling. Move the `smul`/`vadd` sections higher to allow for that spelling. From LeanCamCombi --- .../Algebra/Group/Pointwise/Finset/Basic.lean | 385 ++++++++---------- Mathlib/Algebra/MvPolynomial/Basic.lean | 2 +- 2 files changed, 162 insertions(+), 225 deletions(-) diff --git a/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean b/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean index 22ddb26282b4f..6d68e3fea5280 100644 --- a/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean +++ b/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean @@ -286,6 +286,164 @@ lemma inv_inter (s t : Finset α) : (s ∩ t)⁻¹ = s⁻¹ ∩ t⁻¹ := coe_in end InvolutiveInv +/-! ### Scalar addition/multiplication of finsets -/ + +section SMul +variable [DecidableEq β] [SMul α β] {s s₁ s₂ : Finset α} {t t₁ t₂ u : Finset β} {a : α} {b : β} + +/-- The pointwise product of two finsets `s` and `t`: `s • t = {x • y | x ∈ s, y ∈ t}`. -/ +@[to_additive "The pointwise sum of two finsets `s` and `t`: `s +ᵥ t = {x +ᵥ y | x ∈ s, y ∈ t}`."] +protected def smul : SMul (Finset α) (Finset β) := ⟨image₂ (· • ·)⟩ + +scoped[Pointwise] attribute [instance] Finset.smul Finset.vadd + +@[to_additive] lemma smul_def : s • t = (s ×ˢ t).image fun p : α × β => p.1 • p.2 := rfl + +@[to_additive] +lemma image_smul_product : ((s ×ˢ t).image fun x : α × β => x.fst • x.snd) = s • t := rfl + +@[to_additive] lemma mem_smul {x : β} : x ∈ s • t ↔ ∃ y ∈ s, ∃ z ∈ t, y • z = x := mem_image₂ + +@[to_additive (attr := simp, norm_cast)] +lemma coe_smul (s : Finset α) (t : Finset β) : ↑(s • t) = (s : Set α) • (t : Set β) := coe_image₂ .. + +@[to_additive] lemma smul_mem_smul : a ∈ s → b ∈ t → a • b ∈ s • t := mem_image₂_of_mem + +@[to_additive] lemma smul_card_le : (s • t).card ≤ s.card • t.card := card_image₂_le .. + +@[to_additive (attr := simp)] +lemma empty_smul (t : Finset β) : (∅ : Finset α) • t = ∅ := image₂_empty_left + +@[to_additive (attr := simp)] +lemma smul_empty (s : Finset α) : s • (∅ : Finset β) = ∅ := image₂_empty_right + +@[to_additive (attr := simp)] +lemma smul_eq_empty : s • t = ∅ ↔ s = ∅ ∨ t = ∅ := image₂_eq_empty_iff + +@[to_additive (attr := simp)] +lemma smul_nonempty_iff : (s • t).Nonempty ↔ s.Nonempty ∧ t.Nonempty := image₂_nonempty_iff + +@[to_additive (attr := aesop safe apply (rule_sets := [finsetNonempty]))] +lemma Nonempty.smul : s.Nonempty → t.Nonempty → (s • t).Nonempty := .image₂ + +@[to_additive] lemma Nonempty.of_smul_left : (s • t).Nonempty → s.Nonempty := .of_image₂_left +@[to_additive] lemma Nonempty.of_smul_right : (s • t).Nonempty → t.Nonempty := .of_image₂_right + +@[to_additive] +lemma smul_singleton (b : β) : s • ({b} : Finset β) = s.image (· • b) := image₂_singleton_right + +@[to_additive] +lemma singleton_smul_singleton (a : α) (b : β) : ({a} : Finset α) • ({b} : Finset β) = {a • b} := + image₂_singleton + +@[to_additive (attr := mono, gcongr)] +lemma smul_subset_smul : s₁ ⊆ s₂ → t₁ ⊆ t₂ → s₁ • t₁ ⊆ s₂ • t₂ := image₂_subset + +@[to_additive] lemma smul_subset_smul_left : t₁ ⊆ t₂ → s • t₁ ⊆ s • t₂ := image₂_subset_left +@[to_additive] lemma smul_subset_smul_right : s₁ ⊆ s₂ → s₁ • t ⊆ s₂ • t := image₂_subset_right +@[to_additive] lemma smul_subset_iff : s • t ⊆ u ↔ ∀ a ∈ s, ∀ b ∈ t, a • b ∈ u := image₂_subset_iff + +@[to_additive] +lemma union_smul [DecidableEq α] : (s₁ ∪ s₂) • t = s₁ • t ∪ s₂ • t := image₂_union_left + +@[to_additive] +lemma smul_union : s • (t₁ ∪ t₂) = s • t₁ ∪ s • t₂ := image₂_union_right + +@[to_additive] +lemma inter_smul_subset [DecidableEq α] : (s₁ ∩ s₂) • t ⊆ s₁ • t ∩ s₂ • t := + image₂_inter_subset_left + +@[to_additive] +lemma smul_inter_subset : s • (t₁ ∩ t₂) ⊆ s • t₁ ∩ s • t₂ := image₂_inter_subset_right + +@[to_additive] +lemma inter_smul_union_subset_union [DecidableEq α] : (s₁ ∩ s₂) • (t₁ ∪ t₂) ⊆ s₁ • t₁ ∪ s₂ • t₂ := + image₂_inter_union_subset_union + +@[to_additive] +lemma union_smul_inter_subset_union [DecidableEq α] : (s₁ ∪ s₂) • (t₁ ∩ t₂) ⊆ s₁ • t₁ ∪ s₂ • t₂ := + image₂_union_inter_subset_union + +/-- If a finset `u` is contained in the scalar product of two sets `s • t`, we can find two finsets +`s'`, `t'` such that `s' ⊆ s`, `t' ⊆ t` and `u ⊆ s' • t'`. -/ +@[to_additive +"If a finset `u` is contained in the scalar sum of two sets `s +ᵥ t`, we can find two +finsets `s'`, `t'` such that `s' ⊆ s`, `t' ⊆ t` and `u ⊆ s' +ᵥ t'`."] +lemma subset_smul {s : Set α} {t : Set β} : + ↑u ⊆ s • t → ∃ (s' : Finset α) (t' : Finset β), ↑s' ⊆ s ∧ ↑t' ⊆ t ∧ u ⊆ s' • t' := + subset_set_image₂ + +end SMul + +/-! ### Translation/scaling of finsets -/ + +section SMul +variable [DecidableEq β] [SMul α β] {s s₁ s₂ t : Finset β} {a : α} {b : β} + +/-- The scaling of a finset `s` by a scalar `a`: `a • s = {a • x | x ∈ s}`. -/ +@[to_additive "The translation of a finset `s` by a vector `a`: `a +ᵥ s = {a +ᵥ x | x ∈ s}`."] +protected def smulFinset : SMul α (Finset β) where smul a := image <| (a • ·) + +scoped[Pointwise] attribute [instance] Finset.smulFinset Finset.vaddFinset + +@[to_additive] lemma smul_finset_def : a • s = s.image (a • ·) := rfl + +@[to_additive] lemma image_smul : s.image (a • ·) = a • s := rfl + +@[to_additive] +lemma mem_smul_finset {x : β} : x ∈ a • s ↔ ∃ y, y ∈ s ∧ a • y = x := by + simp only [Finset.smul_finset_def, and_assoc, mem_image, exists_prop, Prod.exists, mem_product] + +@[to_additive (attr := simp, norm_cast)] +lemma coe_smul_finset (a : α) (s : Finset β) : ↑(a • s) = a • (↑s : Set β) := coe_image + +@[to_additive] lemma smul_mem_smul_finset : b ∈ s → a • b ∈ a • s := mem_image_of_mem _ + +@[to_additive] lemma smul_finset_card_le : (a • s).card ≤ s.card := card_image_le + +@[to_additive (attr := simp)] +lemma smul_finset_empty (a : α) : a • (∅ : Finset β) = ∅ := image_empty _ + +@[to_additive (attr := simp)] +lemma smul_finset_eq_empty : a • s = ∅ ↔ s = ∅ := image_eq_empty + +@[to_additive (attr := simp)] +lemma smul_finset_nonempty : (a • s).Nonempty ↔ s.Nonempty := image_nonempty + +@[to_additive (attr := aesop safe apply (rule_sets := [finsetNonempty]))] +lemma Nonempty.smul_finset (hs : s.Nonempty) : (a • s).Nonempty := + hs.image _ + +@[to_additive (attr := simp)] +lemma singleton_smul (a : α) : ({a} : Finset α) • t = a • t := image₂_singleton_left + +@[to_additive (attr := mono, gcongr)] +lemma smul_finset_subset_smul_finset : s ⊆ t → a • s ⊆ a • t := image_subset_image + +@[to_additive (attr := simp)] +lemma smul_finset_singleton (b : β) : a • ({b} : Finset β) = {a • b} := image_singleton .. + +@[to_additive] +lemma smul_finset_union : a • (s₁ ∪ s₂) = a • s₁ ∪ a • s₂ := image_union _ _ + +@[to_additive] +lemma smul_finset_insert (a : α) (b : β) (s : Finset β) : a • insert b s = insert (a • b) (a • s) := + image_insert .. + +@[to_additive] +lemma smul_finset_inter_subset : a • (s₁ ∩ s₂) ⊆ a • s₁ ∩ a • s₂ := image_inter_subset _ _ _ + +@[to_additive] +lemma smul_finset_subset_smul {s : Finset α} : a ∈ s → a • t ⊆ s • t := image_subset_image₂_right + +@[to_additive (attr := simp)] +lemma biUnion_smul_finset (s : Finset α) (t : Finset β) : s.biUnion (· • t) = s • t := + biUnion_image_left + +end SMul + +open scoped Pointwise + /-! ### Finset addition/multiplication -/ @@ -360,13 +518,9 @@ theorem Nonempty.of_mul_left : (s * t).Nonempty → s.Nonempty := theorem Nonempty.of_mul_right : (s * t).Nonempty → t.Nonempty := Nonempty.of_image₂_right -@[to_additive] -theorem mul_singleton (a : α) : s * {a} = s.image (· * a) := - image₂_singleton_right - -@[to_additive] -theorem singleton_mul (a : α) : {a} * s = s.image (a * ·) := - image₂_singleton_left +open scoped RightActions in +@[to_additive] lemma mul_singleton (a : α) : s * {a} = s <• a := image₂_singleton_right +@[to_additive] lemma singleton_mul (a : α) : {a} * s = a • s := image₂_singleton_left @[to_additive (attr := simp)] theorem singleton_mul_singleton (a b : α) : ({a} : Finset α) * {b} = {a * b} := @@ -1011,131 +1165,6 @@ theorem preimage_mul_right_one' : preimage 1 (· * b⁻¹) (mul_left_injective _ end Group -/-! ### Scalar addition/multiplication of finsets -/ - - -section SMul - -variable [DecidableEq β] [SMul α β] {s s₁ s₂ : Finset α} {t t₁ t₂ u : Finset β} {a : α} {b : β} - -/-- The pointwise product of two finsets `s` and `t`: `s • t = {x • y | x ∈ s, y ∈ t}`. -/ -@[to_additive "The pointwise sum of two finsets `s` and `t`: `s +ᵥ t = {x +ᵥ y | x ∈ s, y ∈ t}`."] -protected def smul : SMul (Finset α) (Finset β) := - ⟨image₂ (· • ·)⟩ - -scoped[Pointwise] attribute [instance] Finset.smul Finset.vadd - -@[to_additive] -theorem smul_def : s • t = (s ×ˢ t).image fun p : α × β => p.1 • p.2 := - rfl - -@[to_additive] -theorem image_smul_product : ((s ×ˢ t).image fun x : α × β => x.fst • x.snd) = s • t := - rfl - -@[to_additive] -theorem mem_smul {x : β} : x ∈ s • t ↔ ∃ y ∈ s, ∃ z ∈ t, y • z = x := - mem_image₂ - -@[to_additive (attr := simp, norm_cast)] -theorem coe_smul (s : Finset α) (t : Finset β) : ↑(s • t) = (s : Set α) • (t : Set β) := - coe_image₂ _ _ _ - -@[to_additive] -theorem smul_mem_smul : a ∈ s → b ∈ t → a • b ∈ s • t := - mem_image₂_of_mem - -@[to_additive] -theorem smul_card_le : (s • t).card ≤ s.card • t.card := - card_image₂_le _ _ _ - -@[to_additive (attr := simp)] -theorem empty_smul (t : Finset β) : (∅ : Finset α) • t = ∅ := - image₂_empty_left - -@[to_additive (attr := simp)] -theorem smul_empty (s : Finset α) : s • (∅ : Finset β) = ∅ := - image₂_empty_right - -@[to_additive (attr := simp)] -theorem smul_eq_empty : s • t = ∅ ↔ s = ∅ ∨ t = ∅ := - image₂_eq_empty_iff - -@[to_additive (attr := simp)] -theorem smul_nonempty_iff : (s • t).Nonempty ↔ s.Nonempty ∧ t.Nonempty := - image₂_nonempty_iff - -@[to_additive (attr := aesop safe apply (rule_sets := [finsetNonempty]))] -theorem Nonempty.smul : s.Nonempty → t.Nonempty → (s • t).Nonempty := - Nonempty.image₂ - -@[to_additive] -theorem Nonempty.of_smul_left : (s • t).Nonempty → s.Nonempty := - Nonempty.of_image₂_left - -@[to_additive] -theorem Nonempty.of_smul_right : (s • t).Nonempty → t.Nonempty := - Nonempty.of_image₂_right - -@[to_additive] -theorem smul_singleton (b : β) : s • ({b} : Finset β) = s.image (· • b) := - image₂_singleton_right - -@[to_additive] -theorem singleton_smul_singleton (a : α) (b : β) : ({a} : Finset α) • ({b} : Finset β) = {a • b} := - image₂_singleton - -@[to_additive (attr := mono, gcongr)] -theorem smul_subset_smul : s₁ ⊆ s₂ → t₁ ⊆ t₂ → s₁ • t₁ ⊆ s₂ • t₂ := - image₂_subset - -@[to_additive] -theorem smul_subset_smul_left : t₁ ⊆ t₂ → s • t₁ ⊆ s • t₂ := - image₂_subset_left - -@[to_additive] -theorem smul_subset_smul_right : s₁ ⊆ s₂ → s₁ • t ⊆ s₂ • t := - image₂_subset_right - -@[to_additive] -theorem smul_subset_iff : s • t ⊆ u ↔ ∀ a ∈ s, ∀ b ∈ t, a • b ∈ u := - image₂_subset_iff - -@[to_additive] -theorem union_smul [DecidableEq α] : (s₁ ∪ s₂) • t = s₁ • t ∪ s₂ • t := - image₂_union_left - -@[to_additive] -theorem smul_union : s • (t₁ ∪ t₂) = s • t₁ ∪ s • t₂ := - image₂_union_right - -@[to_additive] -theorem inter_smul_subset [DecidableEq α] : (s₁ ∩ s₂) • t ⊆ s₁ • t ∩ s₂ • t := - image₂_inter_subset_left - -@[to_additive] -theorem smul_inter_subset : s • (t₁ ∩ t₂) ⊆ s • t₁ ∩ s • t₂ := - image₂_inter_subset_right - -@[to_additive] -theorem inter_smul_union_subset_union [DecidableEq α] : (s₁ ∩ s₂) • (t₁ ∪ t₂) ⊆ s₁ • t₁ ∪ s₂ • t₂ := - image₂_inter_union_subset_union - -@[to_additive] -theorem union_smul_inter_subset_union [DecidableEq α] : (s₁ ∪ s₂) • (t₁ ∩ t₂) ⊆ s₁ • t₁ ∪ s₂ • t₂ := - image₂_union_inter_subset_union - -/-- If a finset `u` is contained in the scalar product of two sets `s • t`, we can find two finsets -`s'`, `t'` such that `s' ⊆ s`, `t' ⊆ t` and `u ⊆ s' • t'`. -/ -@[to_additive - "If a finset `u` is contained in the scalar sum of two sets `s +ᵥ t`, we can find two - 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_set_image₂ - -end SMul - /-! ### Scalar subtraction of finsets -/ @@ -1246,98 +1275,6 @@ theorem subset_vsub {s t : Set β} : end VSub -open Pointwise - -/-! ### Translation/scaling of finsets -/ - - -section SMul - -variable [DecidableEq β] [SMul α β] {s s₁ s₂ t : Finset β} {a : α} {b : β} - -/-- The scaling of a finset `s` by a scalar `a`: `a • s = {a • x | x ∈ s}`. -/ -@[to_additive "The translation of a finset `s` by a vector `a`: `a +ᵥ s = {a +ᵥ x | x ∈ s}`."] -protected def smulFinset : SMul α (Finset β) := - ⟨fun a => image <| (a • ·)⟩ - -scoped[Pointwise] attribute [instance] Finset.smulFinset Finset.vaddFinset - -@[to_additive] -theorem smul_finset_def : a • s = s.image (a • ·) := - rfl - -@[to_additive] -theorem image_smul : (s.image fun x => a • x) = a • s := - rfl - -@[to_additive] -theorem mem_smul_finset {x : β} : x ∈ a • s ↔ ∃ y, y ∈ s ∧ a • y = x := by - simp only [Finset.smul_finset_def, and_assoc, mem_image, exists_prop, Prod.exists, mem_product] - -@[to_additive (attr := simp, norm_cast)] -theorem coe_smul_finset (a : α) (s : Finset β) : ↑(a • s) = a • (↑s : Set β) := - coe_image - -@[to_additive] -theorem smul_mem_smul_finset : b ∈ s → a • b ∈ a • s := - mem_image_of_mem _ - -@[to_additive] -theorem smul_finset_card_le : (a • s).card ≤ s.card := - card_image_le - -@[to_additive (attr := simp)] -theorem smul_finset_empty (a : α) : a • (∅ : Finset β) = ∅ := - image_empty _ - -@[to_additive (attr := simp)] -theorem smul_finset_eq_empty : a • s = ∅ ↔ s = ∅ := - image_eq_empty - -@[to_additive (attr := simp)] -theorem smul_finset_nonempty : (a • s).Nonempty ↔ s.Nonempty := - image_nonempty - -@[to_additive (attr := aesop safe apply (rule_sets := [finsetNonempty]))] -theorem Nonempty.smul_finset (hs : s.Nonempty) : (a • s).Nonempty := - hs.image _ - -@[to_additive (attr := simp)] -theorem singleton_smul (a : α) : ({a} : Finset α) • t = a • t := - image₂_singleton_left - -@[to_additive (attr := mono, gcongr)] -theorem smul_finset_subset_smul_finset : s ⊆ t → a • s ⊆ a • t := - image_subset_image - -@[to_additive (attr := simp)] -theorem smul_finset_singleton (b : β) : a • ({b} : Finset β) = {a • b} := - image_singleton _ _ - -@[to_additive] -theorem smul_finset_union : a • (s₁ ∪ s₂) = a • s₁ ∪ a • s₂ := - image_union _ _ - -@[to_additive] -lemma smul_finset_insert (a : α) (b : β) (s : Finset β) : a • insert b s = insert (a • b) (a • s) := - image_insert .. - -@[to_additive] -theorem smul_finset_inter_subset : a • (s₁ ∩ s₂) ⊆ a • s₁ ∩ a • s₂ := - image_inter_subset _ _ _ - -@[to_additive] -theorem smul_finset_subset_smul {s : Finset α} : a ∈ s → a • t ⊆ s • t := - image_subset_image₂_right - -@[to_additive (attr := simp)] -theorem biUnion_smul_finset (s : Finset α) (t : Finset β) : s.biUnion (· • t) = s • t := - biUnion_image_left - -end SMul - -open Pointwise - section Instances variable [DecidableEq γ] diff --git a/Mathlib/Algebra/MvPolynomial/Basic.lean b/Mathlib/Algebra/MvPolynomial/Basic.lean index de4f2a12c09bb..910e409090c04 100644 --- a/Mathlib/Algebra/MvPolynomial/Basic.lean +++ b/Mathlib/Algebra/MvPolynomial/Basic.lean @@ -689,7 +689,7 @@ theorem coeff_mul_monomial' (m) (s : σ →₀ ℕ) (r : R) (p : MvPolynomial σ · contrapose! h rw [← mem_support_iff] at h obtain ⟨j, -, rfl⟩ : ∃ j ∈ support p, j + s = m := by - simpa [Finset.add_singleton] + simpa [Finset.mem_add] using Finset.add_subset_add_left support_monomial_subset <| support_mul _ _ h exact le_add_left le_rfl From 9883fb18a5151b45bfcdefbe61f1d79fc13ef3a8 Mon Sep 17 00:00:00 2001 From: mathlib4-update-dependencies-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Tue, 22 Oct 2024 14:15:56 +0000 Subject: [PATCH 278/505] chore: update Mathlib dependencies 2024-10-22 (#18070) This PR updates the Mathlib dependencies. --- lake-manifest.json | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/lake-manifest.json b/lake-manifest.json index 89a2da4e5b414..a34431f99113a 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,12 +5,12 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "7c5548eeeb1748da5c2872dbd866d3acbaea4b3f", + "rev": "dc72dcdb8e97b3c56bd70f06f043ed2dee3258e6", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", "inherited": false, - "configFile": "lakefile.lean"}, + "configFile": "lakefile.toml"}, {"url": "https://github.com/leanprover-community/quote4", "type": "git", "subDir": null, From f8c840f2ba3c66470aa24aec11803688f2ad4c3f Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Tue, 22 Oct 2024 14:55:23 +0000 Subject: [PATCH 279/505] refactor(Order/ConditionallyCompleteLattice): split large file (#18029) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Violeta Hernández --- Mathlib.lean | 1 + .../Algebra/Order/Group/CompleteLattice.lean | 2 +- Mathlib/Algebra/Order/Group/Indicator.lean | 2 +- .../Group/Pointwise/CompleteLattice.lean | 2 +- Mathlib/Data/Real/Archimedean.lean | 1 + Mathlib/Order/Atoms.lean | 1 + .../ConditionallyCompleteLattice/Basic.lean | 545 +--------------- .../ConditionallyCompleteLattice/Finset.lean | 2 +- .../ConditionallyCompleteLattice/Group.lean | 2 +- .../ConditionallyCompleteLattice/Indexed.lean | 596 ++++++++++++++++++ Mathlib/Order/Filter/AtTopBot.lean | 2 +- Mathlib/Order/Filter/Extr.lean | 2 +- Mathlib/Order/PartialSups.lean | 2 +- Mathlib/Order/SemiconjSup.lean | 1 - .../Order/SuccPred/CompleteLinearOrder.lean | 1 - Mathlib/SetTheory/Cardinal/Basic.lean | 5 +- 16 files changed, 613 insertions(+), 554 deletions(-) create mode 100644 Mathlib/Order/ConditionallyCompleteLattice/Indexed.lean diff --git a/Mathlib.lean b/Mathlib.lean index a904678156efa..0f583235c97b4 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3673,6 +3673,7 @@ import Mathlib.Order.Concept import Mathlib.Order.ConditionallyCompleteLattice.Basic import Mathlib.Order.ConditionallyCompleteLattice.Finset import Mathlib.Order.ConditionallyCompleteLattice.Group +import Mathlib.Order.ConditionallyCompleteLattice.Indexed import Mathlib.Order.Copy import Mathlib.Order.CountableDenseLinearOrder import Mathlib.Order.Cover diff --git a/Mathlib/Algebra/Order/Group/CompleteLattice.lean b/Mathlib/Algebra/Order/Group/CompleteLattice.lean index 19c8f5fda8e68..78a88b90761f3 100644 --- a/Mathlib/Algebra/Order/Group/CompleteLattice.lean +++ b/Mathlib/Algebra/Order/Group/CompleteLattice.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury G. Kudryashov -/ import Mathlib.Algebra.Order.Group.OrderIso -import Mathlib.Order.ConditionallyCompleteLattice.Basic +import Mathlib.Order.ConditionallyCompleteLattice.Indexed /-! # Distributivity of group operations over supremum/infimum diff --git a/Mathlib/Algebra/Order/Group/Indicator.lean b/Mathlib/Algebra/Order/Group/Indicator.lean index a8e229c800374..cc07de02c0ed9 100644 --- a/Mathlib/Algebra/Order/Group/Indicator.lean +++ b/Mathlib/Algebra/Order/Group/Indicator.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import Mathlib.Algebra.Group.Indicator -import Mathlib.Order.ConditionallyCompleteLattice.Basic +import Mathlib.Order.ConditionallyCompleteLattice.Indexed import Mathlib.Algebra.Order.Group.Defs import Mathlib.Algebra.Order.Group.Synonym import Mathlib.Algebra.Order.Group.Unbundled.Abs diff --git a/Mathlib/Algebra/Order/Group/Pointwise/CompleteLattice.lean b/Mathlib/Algebra/Order/Group/Pointwise/CompleteLattice.lean index 6f09d373cffd4..f2539b43a96df 100644 --- a/Mathlib/Algebra/Order/Group/Pointwise/CompleteLattice.lean +++ b/Mathlib/Algebra/Order/Group/Pointwise/CompleteLattice.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import Mathlib.Algebra.Order.Group.Pointwise.Bounds -import Mathlib.Order.ConditionallyCompleteLattice.Basic +import Mathlib.Order.ConditionallyCompleteLattice.Indexed /-! # Infima/suprema in ordered monoids and groups diff --git a/Mathlib/Data/Real/Archimedean.lean b/Mathlib/Data/Real/Archimedean.lean index 5c7543769551c..2e6e3bb63eb7a 100644 --- a/Mathlib/Data/Real/Archimedean.lean +++ b/Mathlib/Data/Real/Archimedean.lean @@ -6,6 +6,7 @@ Authors: Mario Carneiro, Floris van Doorn import Mathlib.Algebra.Order.Archimedean.Basic import Mathlib.Algebra.Order.Group.Pointwise.Bounds import Mathlib.Data.Real.Basic +import Mathlib.Order.ConditionallyCompleteLattice.Indexed import Mathlib.Order.Interval.Set.Disjoint /-! diff --git a/Mathlib/Order/Atoms.lean b/Mathlib/Order/Atoms.lean index d7d36b09473b6..6d4b89c1b3df3 100644 --- a/Mathlib/Order/Atoms.lean +++ b/Mathlib/Order/Atoms.lean @@ -8,6 +8,7 @@ import Mathlib.Order.ModularLattice import Mathlib.Order.SuccPred.Basic import Mathlib.Order.WellFounded import Mathlib.Tactic.Nontriviality +import Mathlib.Order.ConditionallyCompleteLattice.Indexed /-! # Atoms, Coatoms, and Simple Lattices diff --git a/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean b/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean index 614bad616a8b8..718526c723f5e 100644 --- a/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean +++ b/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean @@ -10,7 +10,7 @@ import Mathlib.Order.Interval.Set.Basic import Mathlib.Data.Set.Lattice /-! -# Theory of conditionally complete lattices. +# Theory of conditionally complete lattices A conditionally complete lattice is a lattice in which every non-empty bounded subset `s` has a least upper bound and a greatest lower bound, denoted below by `sSup s` and `sInf s`. @@ -18,9 +18,8 @@ Typical examples are `ℝ`, `ℕ`, and `ℤ` with their usual orders. The theory is very comparable to the theory of complete lattices, except that suitable boundedness and nonemptiness assumptions have to be added to most statements. -We introduce two predicates `BddAbove` and `BddBelow` to express this boundedness, prove -their basic properties, and then go on to prove most useful properties of `sSup` and `sInf` -in conditionally complete lattices. +We express these using the `BddAbove` and `BddBelow` predicates, which we use to prove +most useful properties of `sSup` and `sInf` in conditionally complete lattices. To differentiate the statements between complete lattices and conditionally complete lattices, we prefix `sInf` and `sSup` in the statements by `c`, giving `csInf` and `csSup`. @@ -82,10 +81,6 @@ theorem WithBot.sSup_eq [SupSet α] {s : Set (WithBot α)} (hs : ¬s ⊆ {⊥}) theorem WithTop.sInf_empty [InfSet α] : sInf (∅ : Set (WithTop α)) = ⊤ := if_pos <| by simp -@[simp] -theorem WithTop.iInf_empty [IsEmpty ι] [InfSet α] (f : ι → WithTop α) : - ⨅ i, f i = ⊤ := by rw [iInf, range_eq_empty, WithTop.sInf_empty] - theorem WithTop.coe_sInf' [InfSet α] {s : Set α} (hs : s.Nonempty) (h's : BddBelow s) : ↑(sInf s) = (sInf ((fun (a : α) ↦ ↑a) '' s) : WithTop α) := by classical @@ -98,11 +93,6 @@ theorem WithTop.coe_sInf' [InfSet α] {s : Set α} (hs : s.Nonempty) (h's : BddB · rw [preimage_image_eq] exact Option.some_injective _ -@[norm_cast] -theorem WithTop.coe_iInf [Nonempty ι] [InfSet α] {f : ι → α} (hf : BddBelow (range f)) : - ↑(⨅ i, f i) = (⨅ i, f i : WithTop α) := by - rw [iInf, iInf, WithTop.coe_sInf' (range_nonempty f) hf, ← range_comp, Function.comp_def] - theorem WithTop.coe_sSup' [SupSet α] {s : Set α} (hs : BddAbove s) : ↑(sSup s) = (sSup ((fun (a : α) ↦ ↑a) '' s) : WithTop α) := by classical @@ -111,42 +101,22 @@ theorem WithTop.coe_sSup' [SupSet α] {s : Set α} (hs : BddAbove s) : · exact Option.some_injective _ · rintro ⟨x, _, ⟨⟩⟩ -@[norm_cast] -theorem WithTop.coe_iSup [SupSet α] (f : ι → α) (h : BddAbove (Set.range f)) : - ↑(⨆ i, f i) = (⨆ i, f i : WithTop α) := by - rw [iSup, iSup, WithTop.coe_sSup' h, ← range_comp, Function.comp_def] - @[simp] theorem WithBot.sSup_empty [SupSet α] : sSup (∅ : Set (WithBot α)) = ⊥ := WithTop.sInf_empty (α := αᵒᵈ) @[deprecated (since := "2024-06-10")] alias WithBot.csSup_empty := WithBot.sSup_empty -@[simp] -theorem WithBot.ciSup_empty [IsEmpty ι] [SupSet α] (f : ι → WithBot α) : - ⨆ i, f i = ⊥ := - WithTop.iInf_empty (α := αᵒᵈ) _ - @[norm_cast] theorem WithBot.coe_sSup' [SupSet α] {s : Set α} (hs : s.Nonempty) (h's : BddAbove s) : ↑(sSup s) = (sSup ((fun (a : α) ↦ ↑a) '' s) : WithBot α) := WithTop.coe_sInf' (α := αᵒᵈ) hs h's -@[norm_cast] -theorem WithBot.coe_iSup [Nonempty ι] [SupSet α] {f : ι → α} (hf : BddAbove (range f)) : - ↑(⨆ i, f i) = (⨆ i, f i : WithBot α) := - WithTop.coe_iInf (α := αᵒᵈ) hf - @[norm_cast] theorem WithBot.coe_sInf' [InfSet α] {s : Set α} (hs : BddBelow s) : ↑(sInf s) = (sInf ((fun (a : α) ↦ ↑a) '' s) : WithBot α) := WithTop.coe_sSup' (α := αᵒᵈ) hs -@[norm_cast] -theorem WithBot.coe_iInf [InfSet α] (f : ι → α) (h : BddBelow (Set.range f)) : - ↑(⨅ i, f i) = (⨅ i, f i : WithBot α) := - WithTop.coe_iSup (α := αᵒᵈ) _ h - end /-- A conditionally complete lattice is a lattice in which @@ -456,52 +426,12 @@ theorem csInf_le_iff (h : BddBelow s) (hs : s.Nonempty) : sInf s ≤ a ↔ ∀ b theorem isLUB_csSup (ne : s.Nonempty) (H : BddAbove s) : IsLUB s (sSup s) := ⟨fun _ => le_csSup H, fun _ => csSup_le ne⟩ -theorem isLUB_ciSup [Nonempty ι] {f : ι → α} (H : BddAbove (range f)) : - IsLUB (range f) (⨆ i, f i) := - isLUB_csSup (range_nonempty f) H - -theorem isLUB_ciSup_set {f : β → α} {s : Set β} (H : BddAbove (f '' s)) (Hne : s.Nonempty) : - IsLUB (f '' s) (⨆ i : s, f i) := by - rw [← sSup_image'] - exact isLUB_csSup (Hne.image _) H - theorem isGLB_csInf (ne : s.Nonempty) (H : BddBelow s) : IsGLB s (sInf s) := ⟨fun _ => csInf_le H, fun _ => le_csInf ne⟩ -theorem isGLB_ciInf [Nonempty ι] {f : ι → α} (H : BddBelow (range f)) : - IsGLB (range f) (⨅ i, f i) := - isGLB_csInf (range_nonempty f) H - -theorem isGLB_ciInf_set {f : β → α} {s : Set β} (H : BddBelow (f '' s)) (Hne : s.Nonempty) : - IsGLB (f '' s) (⨅ i : s, f i) := - isLUB_ciSup_set (α := αᵒᵈ) H Hne - -theorem ciSup_le_iff [Nonempty ι] {f : ι → α} {a : α} (hf : BddAbove (range f)) : - iSup f ≤ a ↔ ∀ i, f i ≤ a := - (isLUB_le_iff <| isLUB_ciSup hf).trans forall_mem_range - -theorem le_ciInf_iff [Nonempty ι] {f : ι → α} {a : α} (hf : BddBelow (range f)) : - a ≤ iInf f ↔ ∀ i, a ≤ f i := - (le_isGLB_iff <| isGLB_ciInf hf).trans forall_mem_range - -theorem ciSup_set_le_iff {ι : Type*} {s : Set ι} {f : ι → α} {a : α} (hs : s.Nonempty) - (hf : BddAbove (f '' s)) : ⨆ i : s, f i ≤ a ↔ ∀ i ∈ s, f i ≤ a := - (isLUB_le_iff <| isLUB_ciSup_set hf hs).trans forall_mem_image - -theorem le_ciInf_set_iff {ι : Type*} {s : Set ι} {f : ι → α} {a : α} (hs : s.Nonempty) - (hf : BddBelow (f '' s)) : (a ≤ ⨅ i : s, f i) ↔ ∀ i ∈ s, a ≤ f i := - (le_isGLB_iff <| isGLB_ciInf_set hf hs).trans forall_mem_image - theorem IsLUB.csSup_eq (H : IsLUB s a) (ne : s.Nonempty) : sSup s = a := (isLUB_csSup ne ⟨a, H.1⟩).unique H -theorem IsLUB.ciSup_eq [Nonempty ι] {f : ι → α} (H : IsLUB (range f) a) : ⨆ i, f i = a := - H.csSup_eq (range_nonempty f) - -theorem IsLUB.ciSup_set_eq {s : Set β} {f : β → α} (H : IsLUB (f '' s) a) (Hne : s.Nonempty) : - ⨆ i : s, f i = a := - IsLUB.csSup_eq (image_eq_range f s ▸ H) (image_eq_range f s ▸ Hne.image f) - /-- A greatest element of a set is the supremum of this set. -/ theorem IsGreatest.csSup_eq (H : IsGreatest s a) : sSup s = a := H.isLUB.csSup_eq H.nonempty @@ -512,13 +442,6 @@ theorem IsGreatest.csSup_mem (H : IsGreatest s a) : sSup s ∈ s := theorem IsGLB.csInf_eq (H : IsGLB s a) (ne : s.Nonempty) : sInf s = a := (isGLB_csInf ne ⟨a, H.1⟩).unique H -theorem IsGLB.ciInf_eq [Nonempty ι] {f : ι → α} (H : IsGLB (range f) a) : ⨅ i, f i = a := - H.csInf_eq (range_nonempty f) - -theorem IsGLB.ciInf_set_eq {s : Set β} {f : β → α} (H : IsGLB (f '' s) a) (Hne : s.Nonempty) : - ⨅ i : s, f i = a := - IsGLB.csInf_eq (image_eq_range f s ▸ H) (image_eq_range f s ▸ Hne.image f) - /-- A least element of a set is the infimum of this set. -/ theorem IsLeast.csInf_eq (H : IsLeast s a) : sInf s = a := H.isGLB.csInf_eq H.nonempty @@ -700,145 +623,6 @@ theorem csSup_Ioc (h : a < b) : sSup (Ioc a b) = b := theorem csSup_Ioo [DenselyOrdered α] (h : a < b) : sSup (Ioo a b) = b := (isLUB_Ioo h).csSup_eq (nonempty_Ioo.2 h) -/-- The indexed supremum of a function is bounded above by a uniform bound -/ -theorem ciSup_le [Nonempty ι] {f : ι → α} {c : α} (H : ∀ x, f x ≤ c) : iSup f ≤ c := - csSup_le (range_nonempty f) (by rwa [forall_mem_range]) - -/-- The indexed supremum of a function is bounded below by the value taken at one point -/ -theorem le_ciSup {f : ι → α} (H : BddAbove (range f)) (c : ι) : f c ≤ iSup f := - le_csSup H (mem_range_self _) - -theorem le_ciSup_of_le {f : ι → α} (H : BddAbove (range f)) (c : ι) (h : a ≤ f c) : a ≤ iSup f := - le_trans h (le_ciSup H c) - -/-- The indexed suprema of two functions are comparable if the functions are pointwise comparable -/ -theorem ciSup_mono {f g : ι → α} (B : BddAbove (range g)) (H : ∀ x, f x ≤ g x) : - iSup f ≤ iSup g := by - cases isEmpty_or_nonempty ι - · rw [iSup_of_empty', iSup_of_empty'] - · exact ciSup_le fun x => le_ciSup_of_le B x (H x) - -theorem le_ciSup_set {f : β → α} {s : Set β} (H : BddAbove (f '' s)) {c : β} (hc : c ∈ s) : - f c ≤ ⨆ i : s, f i := - (le_csSup H <| mem_image_of_mem f hc).trans_eq sSup_image' - -/-- The indexed infimum of two functions are comparable if the functions are pointwise comparable -/ -theorem ciInf_mono {f g : ι → α} (B : BddBelow (range f)) (H : ∀ x, f x ≤ g x) : iInf f ≤ iInf g := - ciSup_mono (α := αᵒᵈ) B H - -/-- The indexed minimum of a function is bounded below by a uniform lower bound -/ -theorem le_ciInf [Nonempty ι] {f : ι → α} {c : α} (H : ∀ x, c ≤ f x) : c ≤ iInf f := - ciSup_le (α := αᵒᵈ) H - -/-- The indexed infimum of a function is bounded above by the value taken at one point -/ -theorem ciInf_le {f : ι → α} (H : BddBelow (range f)) (c : ι) : iInf f ≤ f c := - le_ciSup (α := αᵒᵈ) H c - -theorem ciInf_le_of_le {f : ι → α} (H : BddBelow (range f)) (c : ι) (h : f c ≤ a) : iInf f ≤ a := - le_ciSup_of_le (α := αᵒᵈ) H c h - -theorem ciInf_set_le {f : β → α} {s : Set β} (H : BddBelow (f '' s)) {c : β} (hc : c ∈ s) : - ⨅ i : s, f i ≤ f c := - le_ciSup_set (α := αᵒᵈ) H hc - -@[simp] -theorem ciSup_const [hι : Nonempty ι] {a : α} : ⨆ _ : ι, a = a := by - rw [iSup, range_const, csSup_singleton] - -@[simp] -theorem ciInf_const [Nonempty ι] {a : α} : ⨅ _ : ι, a = a := - ciSup_const (α := αᵒᵈ) - -@[simp] -theorem ciSup_unique [Unique ι] {s : ι → α} : ⨆ i, s i = s default := by - have : ∀ i, s i = s default := fun i => congr_arg s (Unique.eq_default i) - simp only [this, ciSup_const] - -@[simp] -theorem ciInf_unique [Unique ι] {s : ι → α} : ⨅ i, s i = s default := - ciSup_unique (α := αᵒᵈ) - -theorem ciSup_subsingleton [Subsingleton ι] (i : ι) (s : ι → α) : ⨆ i, s i = s i := - @ciSup_unique α ι _ ⟨⟨i⟩, fun j => Subsingleton.elim j i⟩ _ - -theorem ciInf_subsingleton [Subsingleton ι] (i : ι) (s : ι → α) : ⨅ i, s i = s i := - @ciInf_unique α ι _ ⟨⟨i⟩, fun j => Subsingleton.elim j i⟩ _ - -@[simp] -theorem ciSup_pos {p : Prop} {f : p → α} (hp : p) : ⨆ h : p, f h = f hp := - ciSup_subsingleton hp f - -@[simp] -theorem ciInf_pos {p : Prop} {f : p → α} (hp : p) : ⨅ h : p, f h = f hp := - ciSup_pos (α := αᵒᵈ) hp - -lemma ciSup_neg {p : Prop} {f : p → α} (hp : ¬ p) : - ⨆ (h : p), f h = sSup (∅ : Set α) := by - rw [iSup] - congr - rwa [range_eq_empty_iff, isEmpty_Prop] - -lemma ciInf_neg {p : Prop} {f : p → α} (hp : ¬ p) : - ⨅ (h : p), f h = sInf (∅ : Set α) := - ciSup_neg (α := αᵒᵈ) hp - -lemma ciSup_eq_ite {p : Prop} [Decidable p] {f : p → α} : - (⨆ h : p, f h) = if h : p then f h else sSup (∅ : Set α) := by - by_cases H : p <;> simp [ciSup_neg, H] - -lemma ciInf_eq_ite {p : Prop} [Decidable p] {f : p → α} : - (⨅ h : p, f h) = if h : p then f h else sInf (∅ : Set α) := - ciSup_eq_ite (α := αᵒᵈ) - -theorem cbiSup_eq_of_forall {p : ι → Prop} {f : Subtype p → α} (hp : ∀ i, p i) : - ⨆ (i) (h : p i), f ⟨i, h⟩ = iSup f := by - simp only [hp, ciSup_unique] - simp only [iSup] - congr - apply Subset.antisymm - · rintro - ⟨i, rfl⟩ - simp [hp i] - · rintro - ⟨i, rfl⟩ - simp - -theorem cbiInf_eq_of_forall {p : ι → Prop} {f : Subtype p → α} (hp : ∀ i, p i) : - ⨅ (i) (h : p i), f ⟨i, h⟩ = iInf f := - cbiSup_eq_of_forall (α := αᵒᵈ) hp - -/-- Introduction rule to prove that `b` is the supremum of `f`: it suffices to check that `b` -is larger than `f i` for all `i`, and that this is not the case of any `w exists_range_iff.mpr <| h₂ w hw - --- Porting note: in mathlib3 `by exact` is not needed -/-- Introduction rule to prove that `b` is the infimum of `f`: it suffices to check that `b` -is smaller than `f i` for all `i`, and that this is not the case of any `w>b`. -See `iInf_eq_of_forall_ge_of_forall_gt_exists_lt` for a version in complete lattices. -/ -theorem ciInf_eq_of_forall_ge_of_forall_gt_exists_lt [Nonempty ι] {f : ι → α} (h₁ : ∀ i, b ≤ f i) - (h₂ : ∀ w, b < w → ∃ i, f i < w) : ⨅ i : ι, f i = b := by - exact ciSup_eq_of_forall_le_of_forall_lt_exists_gt (α := αᵒᵈ) (f := ‹_›) ‹_› ‹_› - -/-- **Nested intervals lemma**: if `f` is a monotone sequence, `g` is an antitone sequence, and -`f n ≤ g n` for all `n`, then `⨆ n, f n` belongs to all the intervals `[f n, g n]`. -/ -theorem Monotone.ciSup_mem_iInter_Icc_of_antitone [SemilatticeSup β] {f g : β → α} (hf : Monotone f) - (hg : Antitone g) (h : f ≤ g) : (⨆ n, f n) ∈ ⋂ n, Icc (f n) (g n) := by - refine mem_iInter.2 fun n => ?_ - haveI : Nonempty β := ⟨n⟩ - have : ∀ m, f m ≤ g n := fun m => hf.forall_le_of_antitone hg h m n - exact ⟨le_ciSup ⟨g <| n, forall_mem_range.2 this⟩ _, ciSup_le this⟩ - -/-- Nested intervals lemma: if `[f n, g n]` is an antitone sequence of nonempty -closed intervals, then `⨆ n, f n` belongs to all the intervals `[f n, g n]`. -/ -theorem ciSup_mem_iInter_Icc_of_antitone_Icc [SemilatticeSup β] {f g : β → α} - (h : Antitone fun n => Icc (f n) (g n)) (h' : ∀ n, f n ≤ g n) : - (⨆ n, f n) ∈ ⋂ n, Icc (f n) (g n) := - Monotone.ciSup_mem_iInter_Icc_of_antitone - (fun _ n hmn => ((Icc_subset_Icc_iff (h' n)).1 (h hmn)).1) - (fun _ n hmn => ((Icc_subset_Icc_iff (h' n)).1 (h hmn)).2) h' - /-- Introduction rule to prove that `b` is the supremum of `s`: it suffices to check that 1) `b` is an upper bound 2) every other upper bound `b'` satisfies `b ≤ b'`. -/ @@ -846,113 +630,12 @@ theorem csSup_eq_of_is_forall_le_of_forall_le_imp_ge (hs : s.Nonempty) (h_is_ub (h_b_le_ub : ∀ ub, (∀ a ∈ s, a ≤ ub) → b ≤ ub) : sSup s = b := (csSup_le hs h_is_ub).antisymm ((h_b_le_ub _) fun _ => le_csSup ⟨b, h_is_ub⟩) -lemma Set.Iic_ciInf [Nonempty ι] {f : ι → α} (hf : BddBelow (range f)) : - Iic (⨅ i, f i) = ⋂ i, Iic (f i) := by - apply Subset.antisymm - · rintro x hx - ⟨i, rfl⟩ - exact hx.trans (ciInf_le hf _) - · rintro x hx - apply le_ciInf - simpa using hx - -lemma Set.Ici_ciSup [Nonempty ι] {f : ι → α} (hf : BddAbove (range f)) : - Ici (⨆ i, f i) = ⋂ i, Ici (f i) := - Iic_ciInf (α := αᵒᵈ) hf - lemma sup_eq_top_of_top_mem [OrderTop α] (h : ⊤ ∈ s) : sSup s = ⊤ := top_unique <| le_csSup (OrderTop.bddAbove s) h lemma inf_eq_bot_of_bot_mem [OrderBot α] (h : ⊥ ∈ s) : sInf s = ⊥ := bot_unique <| csInf_le (OrderBot.bddBelow s) h -theorem ciSup_subtype [Nonempty ι] {p : ι → Prop} [Nonempty (Subtype p)] {f : Subtype p → α} - (hf : BddAbove (Set.range f)) (hf' : sSup ∅ ≤ iSup f) : - iSup f = ⨆ (i) (h : p i), f ⟨i, h⟩ := by - classical - refine le_antisymm (ciSup_le ?_) ?_ - · intro ⟨i, h⟩ - have : f ⟨i, h⟩ = (fun i : ι ↦ ⨆ (h : p i), f ⟨i, h⟩) i := by simp [h] - rw [this] - refine le_ciSup (f := (fun i : ι ↦ ⨆ (h : p i), f ⟨i, h⟩)) ?_ i - simp_rw [ciSup_eq_ite] - refine (hf.union (bddAbove_singleton (a := sSup ∅))).mono ?_ - intro - simp only [Set.mem_range, Set.union_singleton, Set.mem_insert_iff, Subtype.exists, - forall_exists_index] - intro b hb - split_ifs at hb - · exact Or.inr ⟨_, _, hb⟩ - · simp_all - · refine ciSup_le fun i ↦ ?_ - simp_rw [ciSup_eq_ite] - split_ifs - · exact le_ciSup hf ?_ - · exact hf' - -theorem ciInf_subtype [Nonempty ι] {p : ι → Prop} [Nonempty (Subtype p)] {f : Subtype p → α} - (hf : BddBelow (Set.range f)) (hf' : iInf f ≤ sInf ∅) : - iInf f = ⨅ (i) (h : p i), f ⟨i, h⟩ := - ciSup_subtype (α := αᵒᵈ) hf hf' - -theorem ciSup_subtype' [Nonempty ι] {p : ι → Prop} [Nonempty (Subtype p)] {f : ∀ i, p i → α} - (hf : BddAbove (Set.range (fun i : Subtype p ↦ f i i.prop))) - (hf' : sSup ∅ ≤ ⨆ (i : Subtype p), f i i.prop) : - ⨆ (i) (h), f i h = ⨆ x : Subtype p, f x x.property := - (ciSup_subtype (f := fun x => f x.val x.property) hf hf').symm - -theorem ciInf_subtype' [Nonempty ι] {p : ι → Prop} [Nonempty (Subtype p)] {f : ∀ i, p i → α} - (hf : BddBelow (Set.range (fun i : Subtype p ↦ f i i.prop))) - (hf' : ⨅ (i : Subtype p), f i i.prop ≤ sInf ∅) : - ⨅ (i) (h), f i h = ⨅ x : Subtype p, f x x.property := - (ciInf_subtype (f := fun x => f x.val x.property) hf hf').symm - -theorem ciSup_subtype'' {ι} [Nonempty ι] {s : Set ι} (hs : s.Nonempty) {f : ι → α} - (hf : BddAbove (Set.range fun i : s ↦ f i)) (hf' : sSup ∅ ≤ ⨆ i : s, f i) : - ⨆ i : s, f i = ⨆ (t : ι) (_ : t ∈ s), f t := - haveI : Nonempty s := Set.Nonempty.to_subtype hs - ciSup_subtype hf hf' - -theorem ciInf_subtype'' {ι} [Nonempty ι] {s : Set ι} (hs : s.Nonempty) {f : ι → α} - (hf : BddBelow (Set.range fun i : s ↦ f i)) (hf' : ⨅ i : s, f i ≤ sInf ∅) : - ⨅ i : s, f i = ⨅ (t : ι) (_ : t ∈ s), f t := - haveI : Nonempty s := Set.Nonempty.to_subtype hs - ciInf_subtype hf hf' - -theorem csSup_image [Nonempty β] {s : Set β} (hs : s.Nonempty) {f : β → α} - (hf : BddAbove (Set.range fun i : s ↦ f i)) (hf' : sSup ∅ ≤ ⨆ i : s, f i) : - sSup (f '' s) = ⨆ a ∈ s, f a := by - rw [← ciSup_subtype'' hs hf hf', iSup, Set.image_eq_range] - -theorem csInf_image [Nonempty β] {s : Set β} (hs : s.Nonempty) {f : β → α} - (hf : BddBelow (Set.range fun i : s ↦ f i)) (hf' : ⨅ i : s, f i ≤ sInf ∅) : - sInf (f '' s) = ⨅ a ∈ s, f a := - csSup_image (α := αᵒᵈ) hs hf hf' - -lemma ciSup_image {α ι ι' : Type*} [ConditionallyCompleteLattice α] [Nonempty ι] [Nonempty ι'] - {s : Set ι} (hs : s.Nonempty) {f : ι → ι'} {g : ι' → α} - (hf : BddAbove (Set.range fun i : s ↦ g (f i))) (hg' : sSup ∅ ≤ ⨆ i : s, g (f i)) : - ⨆ i ∈ (f '' s), g i = ⨆ x ∈ s, g (f x) := by - have hg : BddAbove (Set.range fun i : f '' s ↦ g i) := by - simpa [bddAbove_def] using hf - have hf' : sSup ∅ ≤ ⨆ i : f '' s, g i := by - refine hg'.trans ?_ - have : Nonempty s := Set.Nonempty.to_subtype hs - refine ciSup_le ?_ - intro ⟨i, h⟩ - obtain ⟨t, ht⟩ : ∃ t : f '' s, g t = g (f (Subtype.mk i h)) := by - have : f i ∈ f '' s := Set.mem_image_of_mem _ h - exact ⟨⟨f i, this⟩, by simp [this]⟩ - rw [← ht] - refine le_ciSup_set ?_ t.prop - simpa [bddAbove_def] using hf - rw [← csSup_image (by simpa using hs) hg hf', ← csSup_image hs hf hg', ← Set.image_comp, comp_def] - -lemma ciInf_image {α ι ι' : Type*} [ConditionallyCompleteLattice α] [Nonempty ι] [Nonempty ι'] - {s : Set ι} (hs : s.Nonempty) {f : ι → ι'} {g : ι' → α} - (hf : BddBelow (Set.range fun i : s ↦ g (f i))) (hg' : ⨅ i : s, g (f i) ≤ sInf ∅) : - ⨅ i ∈ (f '' s), g i = ⨅ x ∈ s, g (f x) := - ciSup_image (α := αᵒᵈ) hs hf hg' - end ConditionallyCompleteLattice instance Pi.conditionallyCompleteLattice {ι : Type*} {α : ι → Type*} @@ -979,38 +662,17 @@ theorem exists_lt_of_lt_csSup (hs : s.Nonempty) (hb : b < sSup s) : ∃ a ∈ s, contrapose! hb exact csSup_le hs hb -/-- Indexed version of the above lemma `exists_lt_of_lt_csSup`. -When `b < iSup f`, there is an element `i` such that `b < f i`. --/ -theorem exists_lt_of_lt_ciSup [Nonempty ι] {f : ι → α} (h : b < iSup f) : ∃ i, b < f i := - let ⟨_, ⟨i, rfl⟩, h⟩ := exists_lt_of_lt_csSup (range_nonempty f) h - ⟨i, h⟩ - /-- When `sInf s < b`, there is an element `a` in `s` with `a < b`, if `s` is nonempty and the order is a linear order. -/ theorem exists_lt_of_csInf_lt (hs : s.Nonempty) (hb : sInf s < b) : ∃ a ∈ s, a < b := exists_lt_of_lt_csSup (α := αᵒᵈ) hs hb -/-- Indexed version of the above lemma `exists_lt_of_csInf_lt` -When `iInf f < a`, there is an element `i` such that `f i < a`. --/ -theorem exists_lt_of_ciInf_lt [Nonempty ι] {f : ι → α} (h : iInf f < a) : ∃ i, f i < a := - exists_lt_of_lt_ciSup (α := αᵒᵈ) h - theorem lt_csSup_iff (hb : BddAbove s) (hs : s.Nonempty) : a < sSup s ↔ ∃ b ∈ s, a < b := by simpa only [not_le, not_forall₂, exists_prop] using (csSup_le_iff hb hs (a := a)).not theorem csInf_lt_iff (hb : BddBelow s) (hs : s.Nonempty) : sInf s < a ↔ ∃ b ∈ s, b < a := by simpa only [not_le, not_forall₂, exists_prop] using (le_csInf_iff hb hs).not -theorem lt_ciSup_iff [Nonempty ι] {f : ι → α} (hb : BddAbove (range f)) : - a < iSup f ↔ ∃ i, a < f i := by - simpa only [mem_range, exists_exists_eq_and] using lt_csSup_iff hb (range_nonempty _) - -theorem ciInf_lt_iff [Nonempty ι] {f : ι → α} (hb : BddBelow (range f)) : - iInf f < a ↔ ∃ i, f i < a := by - simpa only [mem_range, exists_exists_eq_and] using csInf_lt_iff hb (range_nonempty _) - theorem csSup_of_not_bddAbove {s : Set α} (hs : ¬BddAbove s) : sSup s = sSup ∅ := ConditionallyCompleteLinearOrder.csSup_of_not_bddAbove s hs @@ -1075,63 +737,12 @@ lemma sSup_iUnion_Iic (f : ι → α) : sSup (⋃ (i : ι), Iic (f i)) = ⨆ i, lemma sInf_iUnion_Ici (f : ι → α) : sInf (⋃ (i : ι), Ici (f i)) = ⨅ i, f i := sSup_iUnion_Iic (α := αᵒᵈ) f -theorem cbiSup_eq_of_not_forall {p : ι → Prop} {f : Subtype p → α} (hp : ¬ (∀ i, p i)) : - ⨆ (i) (h : p i), f ⟨i, h⟩ = iSup f ⊔ sSup ∅ := by - classical - rcases not_forall.1 hp with ⟨i₀, hi₀⟩ - have : Nonempty ι := ⟨i₀⟩ - simp only [ciSup_eq_ite] - by_cases H : BddAbove (range f) - · have B : BddAbove (range fun i ↦ if h : p i then f ⟨i, h⟩ else sSup ∅) := by - rcases H with ⟨c, hc⟩ - refine ⟨c ⊔ sSup ∅, ?_⟩ - rintro - ⟨i, rfl⟩ - by_cases hi : p i - · simp only [hi, dite_true, le_sup_iff, hc (mem_range_self _), true_or] - · simp only [hi, dite_false, le_sup_right] - apply le_antisymm - · apply ciSup_le (fun i ↦ ?_) - by_cases hi : p i - · simp only [hi, dite_true, le_sup_iff] - left - exact le_ciSup H _ - · simp [hi] - · apply sup_le - · rcases isEmpty_or_nonempty (Subtype p) with hp|hp - · rw [iSup_of_empty'] - convert le_ciSup B i₀ - simp [hi₀] - · apply ciSup_le - rintro ⟨i, hi⟩ - convert le_ciSup B i - simp [hi] - · convert le_ciSup B i₀ - simp [hi₀] - · have : iSup f = sSup (∅ : Set α) := csSup_of_not_bddAbove H - simp only [this, le_refl, sup_of_le_left] - apply csSup_of_not_bddAbove - contrapose! H - apply H.mono - rintro - ⟨i, rfl⟩ - convert mem_range_self i.1 - simp [i.2] - -theorem cbiInf_eq_of_not_forall {p : ι → Prop} {f : Subtype p → α} (hp : ¬ (∀ i, p i)) : - ⨅ (i) (h : p i), f ⟨i, h⟩ = iInf f ⊓ sInf ∅ := - cbiSup_eq_of_not_forall (α := αᵒᵈ) hp - theorem csInf_eq_bot_of_bot_mem [OrderBot α] {s : Set α} (hs : ⊥ ∈ s) : sInf s = ⊥ := eq_bot_iff.2 <| csInf_le (OrderBot.bddBelow s) hs -theorem ciInf_eq_bot_of_bot_mem [OrderBot α] {f : ι → α} (hs : ⊥ ∈ range f) : iInf f = ⊥ := - csInf_eq_bot_of_bot_mem hs - theorem csSup_eq_top_of_top_mem [OrderTop α] {s : Set α} (hs : ⊤ ∈ s) : sSup s = ⊤ := csInf_eq_bot_of_bot_mem (α := αᵒᵈ) hs -theorem ciInf_eq_top_of_top_mem [OrderTop α] {f : ι → α} (hs : ⊤ ∈ range f) : iSup f = ⊤ := - csSup_eq_top_of_top_mem hs - open Function variable [WellFoundedLT α] @@ -1149,9 +760,6 @@ theorem le_csInf_iff' (hs : s.Nonempty) : b ≤ sInf s ↔ b ∈ lowerBounds s : theorem csInf_mem (hs : s.Nonempty) : sInf s ∈ s := (isLeast_csInf hs).1 -theorem ciInf_mem [Nonempty ι] (f : ι → α) : iInf f ∈ range f := - csInf_mem (range_nonempty f) - theorem MonotoneOn.map_csInf {β : Type*} [ConditionallyCompleteLattice β] {f : α → β} (hf : MonotoneOn f s) (hs : s.Nonempty) : f (sInf s) = sInf (f '' s) := (hf.map_isLeast (isLeast_csInf hs)).csInf_eq.symm @@ -1181,13 +789,6 @@ variable [ConditionallyCompleteLinearOrderBot α] {s : Set α} {f : ι → α} { theorem csSup_empty : (sSup ∅ : α) = ⊥ := ConditionallyCompleteLinearOrderBot.csSup_empty -@[simp] -theorem ciSup_of_empty [IsEmpty ι] (f : ι → α) : ⨆ i, f i = ⊥ := by - rw [iSup_of_empty', csSup_empty] - -theorem ciSup_false (f : False → α) : ⨆ i, f i = ⊥ := - ciSup_of_empty f - theorem isLUB_csSup' {s : Set α} (hs : BddAbove s) : IsLUB s (sSup s) := by rcases eq_empty_or_nonempty s with (rfl | hne) · simp only [csSup_empty, isLUB_empty] @@ -1210,51 +811,19 @@ theorem le_csSup_iff' {s : Set α} {a : α} (h : BddAbove s) : a ≤ sSup s ↔ ∀ b, b ∈ upperBounds s → a ≤ b := ⟨fun h _ hb => le_trans h (csSup_le' hb), fun hb => hb _ fun _ => le_csSup h⟩ -theorem le_ciSup_iff' {s : ι → α} {a : α} (h : BddAbove (range s)) : - a ≤ iSup s ↔ ∀ b, (∀ i, s i ≤ b) → a ≤ b := by simp [iSup, h, le_csSup_iff', upperBounds] - theorem le_csInf_iff'' {s : Set α} {a : α} (ne : s.Nonempty) : a ≤ sInf s ↔ ∀ b : α, b ∈ s → a ≤ b := le_csInf_iff (OrderBot.bddBelow _) ne -theorem le_ciInf_iff' [Nonempty ι] {f : ι → α} {a : α} : a ≤ iInf f ↔ ∀ i, a ≤ f i := - le_ciInf_iff (OrderBot.bddBelow _) - theorem csInf_le' (h : a ∈ s) : sInf s ≤ a := csInf_le (OrderBot.bddBelow _) h -theorem ciInf_le' (f : ι → α) (i : ι) : iInf f ≤ f i := ciInf_le (OrderBot.bddBelow _) _ - -lemma ciInf_le_of_le' (c : ι) : f c ≤ a → iInf f ≤ a := ciInf_le_of_le (OrderBot.bddBelow _) _ - theorem exists_lt_of_lt_csSup' {s : Set α} {a : α} (h : a < sSup s) : ∃ b ∈ s, a < b := by contrapose! h exact csSup_le' h -/-- In conditionally complete orders with a bottom element, the nonempty condition can be omitted -from `ciSup_le_iff`. -/ -theorem ciSup_le_iff' {f : ι → α} (h : BddAbove (range f)) {a : α} : - ⨆ i, f i ≤ a ↔ ∀ i, f i ≤ a := - (csSup_le_iff' h).trans forall_mem_range - -theorem ciSup_le' {f : ι → α} {a : α} (h : ∀ i, f i ≤ a) : ⨆ i, f i ≤ a := - csSup_le' <| forall_mem_range.2 h - -/-- In conditionally complete orders with a bottom element, the nonempty condition can be omitted -from `lt_ciSup_iff`. -/ -theorem lt_ciSup_iff' {f : ι → α} (h : BddAbove (range f)) : a < iSup f ↔ ∃ i, a < f i := by - simpa only [not_le, not_forall] using (ciSup_le_iff' h).not - -theorem exists_lt_of_lt_ciSup' {f : ι → α} {a : α} (h : a < ⨆ i, f i) : ∃ i, a < f i := by - contrapose! h - exact ciSup_le' h - theorem not_mem_of_lt_csInf' {x : α} {s : Set α} (h : x < sInf s) : x ∉ s := not_mem_of_lt_csInf h (OrderBot.bddBelow s) -theorem ciSup_mono' {ι'} {f : ι → α} {g : ι' → α} (hg : BddAbove (range g)) - (h : ∀ i, ∃ i', f i ≤ g i') : iSup f ≤ iSup g := - ciSup_le' fun i => Exists.elim (h i) (le_ciSup_of_le hg) - theorem csInf_le_csInf' {s t : Set α} (h₁ : t.Nonempty) (h₂ : t ⊆ s) : sInf s ≤ sInf t := csInf_le_csInf (OrderBot.bddBelow s) h₁ h₂ @@ -1264,14 +833,6 @@ theorem csSup_le_csSup' {s t : Set α} (h₁ : BddAbove t) (h₂ : s ⊆ t) : sS exact bot_le · exact csSup_le_csSup h₁ h h₂ -lemma ciSup_or' (p q : Prop) (f : p ∨ q → α) : - ⨆ (h : p ∨ q), f h = (⨆ h : p, f (.inl h)) ⊔ ⨆ h : q, f (.inr h) := by - by_cases hp : p <;> by_cases hq : q - · simp [hp, hq] - · simp [hp, hq] - · simp [hp, hq] - · simp [hp, hq] - end ConditionallyCompleteLinearOrderBot namespace WithTop @@ -1437,83 +998,6 @@ theorem le_csInf_image {s : Set α} (hs : s.Nonempty) {B : α} (hB : B ∈ lower end Monotone -namespace GaloisConnection - -variable [ConditionallyCompleteLattice α] [ConditionallyCompleteLattice β] [Nonempty ι] {l : α → β} - {u : β → α} - -theorem l_csSup (gc : GaloisConnection l u) {s : Set α} (hne : s.Nonempty) (hbdd : BddAbove s) : - l (sSup s) = ⨆ x : s, l x := - Eq.symm <| IsLUB.ciSup_set_eq (gc.isLUB_l_image <| isLUB_csSup hne hbdd) hne - -theorem l_csSup' (gc : GaloisConnection l u) {s : Set α} (hne : s.Nonempty) (hbdd : BddAbove s) : - l (sSup s) = sSup (l '' s) := by rw [gc.l_csSup hne hbdd, sSup_image'] - -theorem l_ciSup (gc : GaloisConnection l u) {f : ι → α} (hf : BddAbove (range f)) : - l (⨆ i, f i) = ⨆ i, l (f i) := by rw [iSup, gc.l_csSup (range_nonempty _) hf, iSup_range'] - -theorem l_ciSup_set (gc : GaloisConnection l u) {s : Set γ} {f : γ → α} (hf : BddAbove (f '' s)) - (hne : s.Nonempty) : l (⨆ i : s, f i) = ⨆ i : s, l (f i) := by - haveI := hne.to_subtype - rw [image_eq_range] at hf - exact gc.l_ciSup hf - -theorem u_csInf (gc : GaloisConnection l u) {s : Set β} (hne : s.Nonempty) (hbdd : BddBelow s) : - u (sInf s) = ⨅ x : s, u x := - gc.dual.l_csSup hne hbdd - -theorem u_csInf' (gc : GaloisConnection l u) {s : Set β} (hne : s.Nonempty) (hbdd : BddBelow s) : - u (sInf s) = sInf (u '' s) := - gc.dual.l_csSup' hne hbdd - -theorem u_ciInf (gc : GaloisConnection l u) {f : ι → β} (hf : BddBelow (range f)) : - u (⨅ i, f i) = ⨅ i, u (f i) := - gc.dual.l_ciSup hf - -theorem u_ciInf_set (gc : GaloisConnection l u) {s : Set γ} {f : γ → β} (hf : BddBelow (f '' s)) - (hne : s.Nonempty) : u (⨅ i : s, f i) = ⨅ i : s, u (f i) := - gc.dual.l_ciSup_set hf hne - -end GaloisConnection - -namespace OrderIso - -variable [ConditionallyCompleteLattice α] [ConditionallyCompleteLattice β] [Nonempty ι] - -theorem map_csSup (e : α ≃o β) {s : Set α} (hne : s.Nonempty) (hbdd : BddAbove s) : - e (sSup s) = ⨆ x : s, e x := - e.to_galoisConnection.l_csSup hne hbdd - -theorem map_csSup' (e : α ≃o β) {s : Set α} (hne : s.Nonempty) (hbdd : BddAbove s) : - e (sSup s) = sSup (e '' s) := - e.to_galoisConnection.l_csSup' hne hbdd - -theorem map_ciSup (e : α ≃o β) {f : ι → α} (hf : BddAbove (range f)) : - e (⨆ i, f i) = ⨆ i, e (f i) := - e.to_galoisConnection.l_ciSup hf - -theorem map_ciSup_set (e : α ≃o β) {s : Set γ} {f : γ → α} (hf : BddAbove (f '' s)) - (hne : s.Nonempty) : e (⨆ i : s, f i) = ⨆ i : s, e (f i) := - e.to_galoisConnection.l_ciSup_set hf hne - -theorem map_csInf (e : α ≃o β) {s : Set α} (hne : s.Nonempty) (hbdd : BddBelow s) : - e (sInf s) = ⨅ x : s, e x := - e.dual.map_csSup hne hbdd - -theorem map_csInf' (e : α ≃o β) {s : Set α} (hne : s.Nonempty) (hbdd : BddBelow s) : - e (sInf s) = sInf (e '' s) := - e.dual.map_csSup' hne hbdd - -theorem map_ciInf (e : α ≃o β) {f : ι → α} (hf : BddBelow (range f)) : - e (⨅ i, f i) = ⨅ i, e (f i) := - e.dual.map_ciSup hf - -theorem map_ciInf_set (e : α ≃o β) {s : Set γ} {f : γ → α} (hf : BddBelow (f '' s)) - (hne : s.Nonempty) : e (⨅ i : s, f i) = ⨅ i : s, e (f i) := - e.dual.map_ciSup_set hf hne - -end OrderIso - /-! ### Supremum/infimum of `Set.image2` @@ -1521,7 +1005,6 @@ A collection of lemmas showing what happens to the suprema/infima of `s` and `t` a binary function whose partial evaluations are lower/upper adjoints of Galois connections. -/ - section variable [ConditionallyCompleteLattice α] [ConditionallyCompleteLattice β] @@ -1669,26 +1152,4 @@ noncomputable instance WithBot.WithTop.completeLinearOrder {α : Type*} [ConditionallyCompleteLinearOrder α] : CompleteLinearOrder (WithBot (WithTop α)) := { completeLattice, linearOrder, LinearOrder.toBiheytingAlgebra with } -namespace WithTop -variable [ConditionallyCompleteLinearOrderBot α] {f : ι → α} - -lemma iSup_coe_eq_top : ⨆ x, (f x : WithTop α) = ⊤ ↔ ¬BddAbove (range f) := by - rw [iSup_eq_top, not_bddAbove_iff] - refine ⟨fun hf r => ?_, fun hf a ha => ?_⟩ - · rcases hf r (WithTop.coe_lt_top r) with ⟨i, hi⟩ - exact ⟨f i, ⟨i, rfl⟩, WithTop.coe_lt_coe.mp hi⟩ - · rcases hf (a.untop ha.ne) with ⟨-, ⟨i, rfl⟩, hi⟩ - exact ⟨i, by simpa only [WithTop.coe_untop _ ha.ne] using WithTop.coe_lt_coe.mpr hi⟩ - -lemma iSup_coe_lt_top : ⨆ x, (f x : WithTop α) < ⊤ ↔ BddAbove (range f) := - lt_top_iff_ne_top.trans iSup_coe_eq_top.not_left - -lemma iInf_coe_eq_top : ⨅ x, (f x : WithTop α) = ⊤ ↔ IsEmpty ι := by simp [isEmpty_iff] - -lemma iInf_coe_lt_top : ⨅ i, (f i : WithTop α) < ⊤ ↔ Nonempty ι := by - rw [lt_top_iff_ne_top, Ne, iInf_coe_eq_top, not_isEmpty_iff] - -end WithTop end WithTopBot - -set_option linter.style.longFile 1800 diff --git a/Mathlib/Order/ConditionallyCompleteLattice/Finset.lean b/Mathlib/Order/ConditionallyCompleteLattice/Finset.lean index c493a2324fc8c..984c9aa53741f 100644 --- a/Mathlib/Order/ConditionallyCompleteLattice/Finset.lean +++ b/Mathlib/Order/ConditionallyCompleteLattice/Finset.lean @@ -3,7 +3,7 @@ Copyright (c) 2018 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ -import Mathlib.Order.ConditionallyCompleteLattice.Basic +import Mathlib.Order.ConditionallyCompleteLattice.Indexed import Mathlib.Data.Set.Finite /-! diff --git a/Mathlib/Order/ConditionallyCompleteLattice/Group.lean b/Mathlib/Order/ConditionallyCompleteLattice/Group.lean index 70a7b16d5447a..a38d8ea9c6490 100644 --- a/Mathlib/Order/ConditionallyCompleteLattice/Group.lean +++ b/Mathlib/Order/ConditionallyCompleteLattice/Group.lean @@ -3,7 +3,7 @@ Copyright (c) 2018 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ -import Mathlib.Order.ConditionallyCompleteLattice.Basic +import Mathlib.Order.ConditionallyCompleteLattice.Indexed import Mathlib.Algebra.Order.Group.Unbundled.Basic import Mathlib.Algebra.Order.Monoid.Unbundled.OrderDual diff --git a/Mathlib/Order/ConditionallyCompleteLattice/Indexed.lean b/Mathlib/Order/ConditionallyCompleteLattice/Indexed.lean new file mode 100644 index 0000000000000..667d696dff1b2 --- /dev/null +++ b/Mathlib/Order/ConditionallyCompleteLattice/Indexed.lean @@ -0,0 +1,596 @@ +/- +Copyright (c) 2018 Sébastian Gouëzel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sébastian Gouëzel +-/ +import Mathlib.Order.ConditionallyCompleteLattice.Basic + +/-! +# Indexed sup / inf in conditionally complete lattices + +This file proves lemmas about `iSup` and `iInf` for functions valued in a conditionally complete, +rather than complete, lattice. We add a prefix `c` to distinguish them from the versions for +complete lattices, giving names `ciSup_xxx` or `ciInf_xxx`. +-/ + +-- Guard against import creep +assert_not_exists Multiset + +open Function OrderDual Set + +variable {α β γ : Type*} {ι : Sort*} + +section + +/-! +Extension of `iSup` and `iInf` from a preorder `α` to `WithTop α` and `WithBot α` +-/ + +variable [Preorder α] + +@[simp] +theorem WithTop.iInf_empty [IsEmpty ι] [InfSet α] (f : ι → WithTop α) : + ⨅ i, f i = ⊤ := by rw [iInf, range_eq_empty, WithTop.sInf_empty] + +@[norm_cast] +theorem WithTop.coe_iInf [Nonempty ι] [InfSet α] {f : ι → α} (hf : BddBelow (range f)) : + ↑(⨅ i, f i) = (⨅ i, f i : WithTop α) := by + rw [iInf, iInf, WithTop.coe_sInf' (range_nonempty f) hf, ← range_comp, Function.comp_def] + +@[norm_cast] +theorem WithTop.coe_iSup [SupSet α] (f : ι → α) (h : BddAbove (Set.range f)) : + ↑(⨆ i, f i) = (⨆ i, f i : WithTop α) := by + rw [iSup, iSup, WithTop.coe_sSup' h, ← range_comp, Function.comp_def] + +@[simp] +theorem WithBot.ciSup_empty [IsEmpty ι] [SupSet α] (f : ι → WithBot α) : + ⨆ i, f i = ⊥ := + WithTop.iInf_empty (α := αᵒᵈ) _ + +@[norm_cast] +theorem WithBot.coe_iSup [Nonempty ι] [SupSet α] {f : ι → α} (hf : BddAbove (range f)) : + ↑(⨆ i, f i) = (⨆ i, f i : WithBot α) := + WithTop.coe_iInf (α := αᵒᵈ) hf + +@[norm_cast] +theorem WithBot.coe_iInf [InfSet α] (f : ι → α) (h : BddBelow (Set.range f)) : + ↑(⨅ i, f i) = (⨅ i, f i : WithBot α) := + WithTop.coe_iSup (α := αᵒᵈ) _ h + +end + +section ConditionallyCompleteLattice + +variable [ConditionallyCompleteLattice α] {s t : Set α} {a b : α} + +theorem isLUB_ciSup [Nonempty ι] {f : ι → α} (H : BddAbove (range f)) : + IsLUB (range f) (⨆ i, f i) := + isLUB_csSup (range_nonempty f) H + +theorem isLUB_ciSup_set {f : β → α} {s : Set β} (H : BddAbove (f '' s)) (Hne : s.Nonempty) : + IsLUB (f '' s) (⨆ i : s, f i) := by + rw [← sSup_image'] + exact isLUB_csSup (Hne.image _) H + +theorem isGLB_ciInf [Nonempty ι] {f : ι → α} (H : BddBelow (range f)) : + IsGLB (range f) (⨅ i, f i) := + isGLB_csInf (range_nonempty f) H + +theorem isGLB_ciInf_set {f : β → α} {s : Set β} (H : BddBelow (f '' s)) (Hne : s.Nonempty) : + IsGLB (f '' s) (⨅ i : s, f i) := + isLUB_ciSup_set (α := αᵒᵈ) H Hne + +theorem ciSup_le_iff [Nonempty ι] {f : ι → α} {a : α} (hf : BddAbove (range f)) : + iSup f ≤ a ↔ ∀ i, f i ≤ a := + (isLUB_le_iff <| isLUB_ciSup hf).trans forall_mem_range + +theorem le_ciInf_iff [Nonempty ι] {f : ι → α} {a : α} (hf : BddBelow (range f)) : + a ≤ iInf f ↔ ∀ i, a ≤ f i := + (le_isGLB_iff <| isGLB_ciInf hf).trans forall_mem_range + +theorem ciSup_set_le_iff {ι : Type*} {s : Set ι} {f : ι → α} {a : α} (hs : s.Nonempty) + (hf : BddAbove (f '' s)) : ⨆ i : s, f i ≤ a ↔ ∀ i ∈ s, f i ≤ a := + (isLUB_le_iff <| isLUB_ciSup_set hf hs).trans forall_mem_image + +theorem le_ciInf_set_iff {ι : Type*} {s : Set ι} {f : ι → α} {a : α} (hs : s.Nonempty) + (hf : BddBelow (f '' s)) : (a ≤ ⨅ i : s, f i) ↔ ∀ i ∈ s, a ≤ f i := + (le_isGLB_iff <| isGLB_ciInf_set hf hs).trans forall_mem_image + +theorem IsLUB.ciSup_eq [Nonempty ι] {f : ι → α} (H : IsLUB (range f) a) : ⨆ i, f i = a := + H.csSup_eq (range_nonempty f) + +theorem IsLUB.ciSup_set_eq {s : Set β} {f : β → α} (H : IsLUB (f '' s) a) (Hne : s.Nonempty) : + ⨆ i : s, f i = a := + IsLUB.csSup_eq (image_eq_range f s ▸ H) (image_eq_range f s ▸ Hne.image f) + +theorem IsGLB.ciInf_eq [Nonempty ι] {f : ι → α} (H : IsGLB (range f) a) : ⨅ i, f i = a := + H.csInf_eq (range_nonempty f) + +theorem IsGLB.ciInf_set_eq {s : Set β} {f : β → α} (H : IsGLB (f '' s) a) (Hne : s.Nonempty) : + ⨅ i : s, f i = a := + IsGLB.csInf_eq (image_eq_range f s ▸ H) (image_eq_range f s ▸ Hne.image f) + +/-- The indexed supremum of a function is bounded above by a uniform bound -/ +theorem ciSup_le [Nonempty ι] {f : ι → α} {c : α} (H : ∀ x, f x ≤ c) : iSup f ≤ c := + csSup_le (range_nonempty f) (by rwa [forall_mem_range]) + +/-- The indexed supremum of a function is bounded below by the value taken at one point -/ +theorem le_ciSup {f : ι → α} (H : BddAbove (range f)) (c : ι) : f c ≤ iSup f := + le_csSup H (mem_range_self _) + +theorem le_ciSup_of_le {f : ι → α} (H : BddAbove (range f)) (c : ι) (h : a ≤ f c) : a ≤ iSup f := + le_trans h (le_ciSup H c) + +/-- The indexed suprema of two functions are comparable if the functions are pointwise comparable -/ +theorem ciSup_mono {f g : ι → α} (B : BddAbove (range g)) (H : ∀ x, f x ≤ g x) : + iSup f ≤ iSup g := by + cases isEmpty_or_nonempty ι + · rw [iSup_of_empty', iSup_of_empty'] + · exact ciSup_le fun x => le_ciSup_of_le B x (H x) + +theorem le_ciSup_set {f : β → α} {s : Set β} (H : BddAbove (f '' s)) {c : β} (hc : c ∈ s) : + f c ≤ ⨆ i : s, f i := + (le_csSup H <| mem_image_of_mem f hc).trans_eq sSup_image' + +/-- The indexed infimum of two functions are comparable if the functions are pointwise comparable -/ +theorem ciInf_mono {f g : ι → α} (B : BddBelow (range f)) (H : ∀ x, f x ≤ g x) : iInf f ≤ iInf g := + ciSup_mono (α := αᵒᵈ) B H + +/-- The indexed minimum of a function is bounded below by a uniform lower bound -/ +theorem le_ciInf [Nonempty ι] {f : ι → α} {c : α} (H : ∀ x, c ≤ f x) : c ≤ iInf f := + ciSup_le (α := αᵒᵈ) H + +/-- The indexed infimum of a function is bounded above by the value taken at one point -/ +theorem ciInf_le {f : ι → α} (H : BddBelow (range f)) (c : ι) : iInf f ≤ f c := + le_ciSup (α := αᵒᵈ) H c + +theorem ciInf_le_of_le {f : ι → α} (H : BddBelow (range f)) (c : ι) (h : f c ≤ a) : iInf f ≤ a := + le_ciSup_of_le (α := αᵒᵈ) H c h + +theorem ciInf_set_le {f : β → α} {s : Set β} (H : BddBelow (f '' s)) {c : β} (hc : c ∈ s) : + ⨅ i : s, f i ≤ f c := + le_ciSup_set (α := αᵒᵈ) H hc + +@[simp] +theorem ciSup_const [hι : Nonempty ι] {a : α} : ⨆ _ : ι, a = a := by + rw [iSup, range_const, csSup_singleton] + +@[simp] +theorem ciInf_const [Nonempty ι] {a : α} : ⨅ _ : ι, a = a := + ciSup_const (α := αᵒᵈ) + +@[simp] +theorem ciSup_unique [Unique ι] {s : ι → α} : ⨆ i, s i = s default := by + have : ∀ i, s i = s default := fun i => congr_arg s (Unique.eq_default i) + simp only [this, ciSup_const] + +@[simp] +theorem ciInf_unique [Unique ι] {s : ι → α} : ⨅ i, s i = s default := + ciSup_unique (α := αᵒᵈ) + +theorem ciSup_subsingleton [Subsingleton ι] (i : ι) (s : ι → α) : ⨆ i, s i = s i := + @ciSup_unique α ι _ ⟨⟨i⟩, fun j => Subsingleton.elim j i⟩ _ + +theorem ciInf_subsingleton [Subsingleton ι] (i : ι) (s : ι → α) : ⨅ i, s i = s i := + @ciInf_unique α ι _ ⟨⟨i⟩, fun j => Subsingleton.elim j i⟩ _ + +@[simp] +theorem ciSup_pos {p : Prop} {f : p → α} (hp : p) : ⨆ h : p, f h = f hp := + ciSup_subsingleton hp f + +@[simp] +theorem ciInf_pos {p : Prop} {f : p → α} (hp : p) : ⨅ h : p, f h = f hp := + ciSup_pos (α := αᵒᵈ) hp + +lemma ciSup_neg {p : Prop} {f : p → α} (hp : ¬ p) : + ⨆ (h : p), f h = sSup (∅ : Set α) := by + rw [iSup] + congr + rwa [range_eq_empty_iff, isEmpty_Prop] + +lemma ciInf_neg {p : Prop} {f : p → α} (hp : ¬ p) : + ⨅ (h : p), f h = sInf (∅ : Set α) := + ciSup_neg (α := αᵒᵈ) hp + +lemma ciSup_eq_ite {p : Prop} [Decidable p] {f : p → α} : + (⨆ h : p, f h) = if h : p then f h else sSup (∅ : Set α) := by + by_cases H : p <;> simp [ciSup_neg, H] + +lemma ciInf_eq_ite {p : Prop} [Decidable p] {f : p → α} : + (⨅ h : p, f h) = if h : p then f h else sInf (∅ : Set α) := + ciSup_eq_ite (α := αᵒᵈ) + +theorem cbiSup_eq_of_forall {p : ι → Prop} {f : Subtype p → α} (hp : ∀ i, p i) : + ⨆ (i) (h : p i), f ⟨i, h⟩ = iSup f := by + simp only [hp, ciSup_unique] + simp only [iSup] + congr + apply Subset.antisymm + · rintro - ⟨i, rfl⟩ + simp [hp i] + · rintro - ⟨i, rfl⟩ + simp + +theorem cbiInf_eq_of_forall {p : ι → Prop} {f : Subtype p → α} (hp : ∀ i, p i) : + ⨅ (i) (h : p i), f ⟨i, h⟩ = iInf f := + cbiSup_eq_of_forall (α := αᵒᵈ) hp + +/-- Introduction rule to prove that `b` is the supremum of `f`: it suffices to check that `b` +is larger than `f i` for all `i`, and that this is not the case of any `w exists_range_iff.mpr <| h₂ w hw + +-- Porting note: in mathlib3 `by exact` is not needed +/-- Introduction rule to prove that `b` is the infimum of `f`: it suffices to check that `b` +is smaller than `f i` for all `i`, and that this is not the case of any `w>b`. +See `iInf_eq_of_forall_ge_of_forall_gt_exists_lt` for a version in complete lattices. -/ +theorem ciInf_eq_of_forall_ge_of_forall_gt_exists_lt [Nonempty ι] {f : ι → α} (h₁ : ∀ i, b ≤ f i) + (h₂ : ∀ w, b < w → ∃ i, f i < w) : ⨅ i : ι, f i = b := by + exact ciSup_eq_of_forall_le_of_forall_lt_exists_gt (α := αᵒᵈ) (f := ‹_›) ‹_› ‹_› + +/-- **Nested intervals lemma**: if `f` is a monotone sequence, `g` is an antitone sequence, and +`f n ≤ g n` for all `n`, then `⨆ n, f n` belongs to all the intervals `[f n, g n]`. -/ +theorem Monotone.ciSup_mem_iInter_Icc_of_antitone [SemilatticeSup β] {f g : β → α} (hf : Monotone f) + (hg : Antitone g) (h : f ≤ g) : (⨆ n, f n) ∈ ⋂ n, Icc (f n) (g n) := by + refine mem_iInter.2 fun n => ?_ + haveI : Nonempty β := ⟨n⟩ + have : ∀ m, f m ≤ g n := fun m => hf.forall_le_of_antitone hg h m n + exact ⟨le_ciSup ⟨g <| n, forall_mem_range.2 this⟩ _, ciSup_le this⟩ + +/-- Nested intervals lemma: if `[f n, g n]` is an antitone sequence of nonempty +closed intervals, then `⨆ n, f n` belongs to all the intervals `[f n, g n]`. -/ +theorem ciSup_mem_iInter_Icc_of_antitone_Icc [SemilatticeSup β] {f g : β → α} + (h : Antitone fun n => Icc (f n) (g n)) (h' : ∀ n, f n ≤ g n) : + (⨆ n, f n) ∈ ⋂ n, Icc (f n) (g n) := + Monotone.ciSup_mem_iInter_Icc_of_antitone + (fun _ n hmn => ((Icc_subset_Icc_iff (h' n)).1 (h hmn)).1) + (fun _ n hmn => ((Icc_subset_Icc_iff (h' n)).1 (h hmn)).2) h' + +lemma Set.Iic_ciInf [Nonempty ι] {f : ι → α} (hf : BddBelow (range f)) : + Iic (⨅ i, f i) = ⋂ i, Iic (f i) := by + apply Subset.antisymm + · rintro x hx - ⟨i, rfl⟩ + exact hx.trans (ciInf_le hf _) + · rintro x hx + apply le_ciInf + simpa using hx + +lemma Set.Ici_ciSup [Nonempty ι] {f : ι → α} (hf : BddAbove (range f)) : + Ici (⨆ i, f i) = ⋂ i, Ici (f i) := + Iic_ciInf (α := αᵒᵈ) hf + +theorem ciSup_subtype [Nonempty ι] {p : ι → Prop} [Nonempty (Subtype p)] {f : Subtype p → α} + (hf : BddAbove (Set.range f)) (hf' : sSup ∅ ≤ iSup f) : + iSup f = ⨆ (i) (h : p i), f ⟨i, h⟩ := by + classical + refine le_antisymm (ciSup_le ?_) ?_ + · intro ⟨i, h⟩ + have : f ⟨i, h⟩ = (fun i : ι ↦ ⨆ (h : p i), f ⟨i, h⟩) i := by simp [h] + rw [this] + refine le_ciSup (f := (fun i : ι ↦ ⨆ (h : p i), f ⟨i, h⟩)) ?_ i + simp_rw [ciSup_eq_ite] + refine (hf.union (bddAbove_singleton (a := sSup ∅))).mono ?_ + intro + simp only [Set.mem_range, Set.union_singleton, Set.mem_insert_iff, Subtype.exists, + forall_exists_index] + intro b hb + split_ifs at hb + · exact Or.inr ⟨_, _, hb⟩ + · simp_all + · refine ciSup_le fun i ↦ ?_ + simp_rw [ciSup_eq_ite] + split_ifs + · exact le_ciSup hf ?_ + · exact hf' + +theorem ciInf_subtype [Nonempty ι] {p : ι → Prop} [Nonempty (Subtype p)] {f : Subtype p → α} + (hf : BddBelow (Set.range f)) (hf' : iInf f ≤ sInf ∅) : + iInf f = ⨅ (i) (h : p i), f ⟨i, h⟩ := + ciSup_subtype (α := αᵒᵈ) hf hf' + +theorem ciSup_subtype' [Nonempty ι] {p : ι → Prop} [Nonempty (Subtype p)] {f : ∀ i, p i → α} + (hf : BddAbove (Set.range (fun i : Subtype p ↦ f i i.prop))) + (hf' : sSup ∅ ≤ ⨆ (i : Subtype p), f i i.prop) : + ⨆ (i) (h), f i h = ⨆ x : Subtype p, f x x.property := + (ciSup_subtype (f := fun x => f x.val x.property) hf hf').symm + +theorem ciInf_subtype' [Nonempty ι] {p : ι → Prop} [Nonempty (Subtype p)] {f : ∀ i, p i → α} + (hf : BddBelow (Set.range (fun i : Subtype p ↦ f i i.prop))) + (hf' : ⨅ (i : Subtype p), f i i.prop ≤ sInf ∅) : + ⨅ (i) (h), f i h = ⨅ x : Subtype p, f x x.property := + (ciInf_subtype (f := fun x => f x.val x.property) hf hf').symm + +theorem ciSup_subtype'' {ι} [Nonempty ι] {s : Set ι} (hs : s.Nonempty) {f : ι → α} + (hf : BddAbove (Set.range fun i : s ↦ f i)) (hf' : sSup ∅ ≤ ⨆ i : s, f i) : + ⨆ i : s, f i = ⨆ (t : ι) (_ : t ∈ s), f t := + haveI : Nonempty s := Set.Nonempty.to_subtype hs + ciSup_subtype hf hf' + +theorem ciInf_subtype'' {ι} [Nonempty ι] {s : Set ι} (hs : s.Nonempty) {f : ι → α} + (hf : BddBelow (Set.range fun i : s ↦ f i)) (hf' : ⨅ i : s, f i ≤ sInf ∅) : + ⨅ i : s, f i = ⨅ (t : ι) (_ : t ∈ s), f t := + haveI : Nonempty s := Set.Nonempty.to_subtype hs + ciInf_subtype hf hf' + +theorem csSup_image [Nonempty β] {s : Set β} (hs : s.Nonempty) {f : β → α} + (hf : BddAbove (Set.range fun i : s ↦ f i)) (hf' : sSup ∅ ≤ ⨆ i : s, f i) : + sSup (f '' s) = ⨆ a ∈ s, f a := by + rw [← ciSup_subtype'' hs hf hf', iSup, Set.image_eq_range] + +theorem csInf_image [Nonempty β] {s : Set β} (hs : s.Nonempty) {f : β → α} + (hf : BddBelow (Set.range fun i : s ↦ f i)) (hf' : ⨅ i : s, f i ≤ sInf ∅) : + sInf (f '' s) = ⨅ a ∈ s, f a := + csSup_image (α := αᵒᵈ) hs hf hf' + +lemma ciSup_image {α ι ι' : Type*} [ConditionallyCompleteLattice α] [Nonempty ι] [Nonempty ι'] + {s : Set ι} (hs : s.Nonempty) {f : ι → ι'} {g : ι' → α} + (hf : BddAbove (Set.range fun i : s ↦ g (f i))) (hg' : sSup ∅ ≤ ⨆ i : s, g (f i)) : + ⨆ i ∈ (f '' s), g i = ⨆ x ∈ s, g (f x) := by + have hg : BddAbove (Set.range fun i : f '' s ↦ g i) := by + simpa [bddAbove_def] using hf + have hf' : sSup ∅ ≤ ⨆ i : f '' s, g i := by + refine hg'.trans ?_ + have : Nonempty s := Set.Nonempty.to_subtype hs + refine ciSup_le ?_ + intro ⟨i, h⟩ + obtain ⟨t, ht⟩ : ∃ t : f '' s, g t = g (f (Subtype.mk i h)) := by + have : f i ∈ f '' s := Set.mem_image_of_mem _ h + exact ⟨⟨f i, this⟩, by simp [this]⟩ + rw [← ht] + refine le_ciSup_set ?_ t.prop + simpa [bddAbove_def] using hf + rw [← csSup_image (by simpa using hs) hg hf', ← csSup_image hs hf hg', ← Set.image_comp, comp_def] + +lemma ciInf_image {α ι ι' : Type*} [ConditionallyCompleteLattice α] [Nonempty ι] [Nonempty ι'] + {s : Set ι} (hs : s.Nonempty) {f : ι → ι'} {g : ι' → α} + (hf : BddBelow (Set.range fun i : s ↦ g (f i))) (hg' : ⨅ i : s, g (f i) ≤ sInf ∅) : + ⨅ i ∈ (f '' s), g i = ⨅ x ∈ s, g (f x) := + ciSup_image (α := αᵒᵈ) hs hf hg' + +end ConditionallyCompleteLattice + +section ConditionallyCompleteLinearOrder + +variable [ConditionallyCompleteLinearOrder α] {s t : Set α} {a b : α} + +/-- Indexed version of `exists_lt_of_lt_csSup`. +When `b < iSup f`, there is an element `i` such that `b < f i`. +-/ +theorem exists_lt_of_lt_ciSup [Nonempty ι] {f : ι → α} (h : b < iSup f) : ∃ i, b < f i := + let ⟨_, ⟨i, rfl⟩, h⟩ := exists_lt_of_lt_csSup (range_nonempty f) h + ⟨i, h⟩ + +/-- Indexed version of `exists_lt_of_csInf_lt`. +When `iInf f < a`, there is an element `i` such that `f i < a`. +-/ +theorem exists_lt_of_ciInf_lt [Nonempty ι] {f : ι → α} (h : iInf f < a) : ∃ i, f i < a := + exists_lt_of_lt_ciSup (α := αᵒᵈ) h + +theorem lt_ciSup_iff [Nonempty ι] {f : ι → α} (hb : BddAbove (range f)) : + a < iSup f ↔ ∃ i, a < f i := by + simpa only [mem_range, exists_exists_eq_and] using lt_csSup_iff hb (range_nonempty _) + +theorem ciInf_lt_iff [Nonempty ι] {f : ι → α} (hb : BddBelow (range f)) : + iInf f < a ↔ ∃ i, f i < a := by + simpa only [mem_range, exists_exists_eq_and] using csInf_lt_iff hb (range_nonempty _) + +theorem cbiSup_eq_of_not_forall {p : ι → Prop} {f : Subtype p → α} (hp : ¬ (∀ i, p i)) : + ⨆ (i) (h : p i), f ⟨i, h⟩ = iSup f ⊔ sSup ∅ := by + classical + rcases not_forall.1 hp with ⟨i₀, hi₀⟩ + have : Nonempty ι := ⟨i₀⟩ + simp only [ciSup_eq_ite] + by_cases H : BddAbove (range f) + · have B : BddAbove (range fun i ↦ if h : p i then f ⟨i, h⟩ else sSup ∅) := by + rcases H with ⟨c, hc⟩ + refine ⟨c ⊔ sSup ∅, ?_⟩ + rintro - ⟨i, rfl⟩ + by_cases hi : p i + · simp only [hi, dite_true, le_sup_iff, hc (mem_range_self _), true_or] + · simp only [hi, dite_false, le_sup_right] + apply le_antisymm + · apply ciSup_le (fun i ↦ ?_) + by_cases hi : p i + · simp only [hi, dite_true, le_sup_iff] + left + exact le_ciSup H _ + · simp [hi] + · apply sup_le + · rcases isEmpty_or_nonempty (Subtype p) with hp|hp + · rw [iSup_of_empty'] + convert le_ciSup B i₀ + simp [hi₀] + · apply ciSup_le + rintro ⟨i, hi⟩ + convert le_ciSup B i + simp [hi] + · convert le_ciSup B i₀ + simp [hi₀] + · have : iSup f = sSup (∅ : Set α) := csSup_of_not_bddAbove H + simp only [this, le_refl, sup_of_le_left] + apply csSup_of_not_bddAbove + contrapose! H + apply H.mono + rintro - ⟨i, rfl⟩ + convert mem_range_self i.1 + simp [i.2] + +theorem cbiInf_eq_of_not_forall {p : ι → Prop} {f : Subtype p → α} (hp : ¬ (∀ i, p i)) : + ⨅ (i) (h : p i), f ⟨i, h⟩ = iInf f ⊓ sInf ∅ := + cbiSup_eq_of_not_forall (α := αᵒᵈ) hp + +theorem ciInf_eq_bot_of_bot_mem [OrderBot α] {f : ι → α} (hs : ⊥ ∈ range f) : iInf f = ⊥ := + csInf_eq_bot_of_bot_mem hs + +theorem ciInf_eq_top_of_top_mem [OrderTop α] {f : ι → α} (hs : ⊤ ∈ range f) : iSup f = ⊤ := + csSup_eq_top_of_top_mem hs + +variable [WellFoundedLT α] + +theorem ciInf_mem [Nonempty ι] (f : ι → α) : iInf f ∈ range f := + csInf_mem (range_nonempty f) + +end ConditionallyCompleteLinearOrder + +/-! +### Lemmas about a conditionally complete linear order with bottom element + +In this case we have `Sup ∅ = ⊥`, so we can drop some `Nonempty`/`Set.Nonempty` assumptions. +-/ + + +section ConditionallyCompleteLinearOrderBot + +variable [ConditionallyCompleteLinearOrderBot α] {s : Set α} {f : ι → α} {a : α} + +@[simp] +theorem ciSup_of_empty [IsEmpty ι] (f : ι → α) : ⨆ i, f i = ⊥ := by + rw [iSup_of_empty', csSup_empty] + +theorem ciSup_false (f : False → α) : ⨆ i, f i = ⊥ := + ciSup_of_empty f + +theorem le_ciSup_iff' {s : ι → α} {a : α} (h : BddAbove (range s)) : + a ≤ iSup s ↔ ∀ b, (∀ i, s i ≤ b) → a ≤ b := by simp [iSup, h, le_csSup_iff', upperBounds] + +theorem le_ciInf_iff' [Nonempty ι] {f : ι → α} {a : α} : a ≤ iInf f ↔ ∀ i, a ≤ f i := + le_ciInf_iff (OrderBot.bddBelow _) + +theorem ciInf_le' (f : ι → α) (i : ι) : iInf f ≤ f i := ciInf_le (OrderBot.bddBelow _) _ + +lemma ciInf_le_of_le' (c : ι) : f c ≤ a → iInf f ≤ a := ciInf_le_of_le (OrderBot.bddBelow _) _ + +/-- In conditionally complete orders with a bottom element, the nonempty condition can be omitted +from `ciSup_le_iff`. -/ +theorem ciSup_le_iff' {f : ι → α} (h : BddAbove (range f)) {a : α} : + ⨆ i, f i ≤ a ↔ ∀ i, f i ≤ a := + (csSup_le_iff' h).trans forall_mem_range + +theorem ciSup_le' {f : ι → α} {a : α} (h : ∀ i, f i ≤ a) : ⨆ i, f i ≤ a := + csSup_le' <| forall_mem_range.2 h + +/-- In conditionally complete orders with a bottom element, the nonempty condition can be omitted +from `lt_ciSup_iff`. -/ +theorem lt_ciSup_iff' {f : ι → α} (h : BddAbove (range f)) : a < iSup f ↔ ∃ i, a < f i := by + simpa only [not_le, not_forall] using (ciSup_le_iff' h).not + +theorem exists_lt_of_lt_ciSup' {f : ι → α} {a : α} (h : a < ⨆ i, f i) : ∃ i, a < f i := by + contrapose! h + exact ciSup_le' h + +theorem ciSup_mono' {ι'} {f : ι → α} {g : ι' → α} (hg : BddAbove (range g)) + (h : ∀ i, ∃ i', f i ≤ g i') : iSup f ≤ iSup g := + ciSup_le' fun i => Exists.elim (h i) (le_ciSup_of_le hg) + +lemma ciSup_or' (p q : Prop) (f : p ∨ q → α) : + ⨆ (h : p ∨ q), f h = (⨆ h : p, f (.inl h)) ⊔ ⨆ h : q, f (.inr h) := by + by_cases hp : p <;> + by_cases hq : q <;> + simp [hp, hq] + +end ConditionallyCompleteLinearOrderBot + +namespace GaloisConnection + +variable [ConditionallyCompleteLattice α] [ConditionallyCompleteLattice β] [Nonempty ι] {l : α → β} + {u : β → α} + +theorem l_csSup (gc : GaloisConnection l u) {s : Set α} (hne : s.Nonempty) (hbdd : BddAbove s) : + l (sSup s) = ⨆ x : s, l x := + Eq.symm <| IsLUB.ciSup_set_eq (gc.isLUB_l_image <| isLUB_csSup hne hbdd) hne + +theorem l_csSup' (gc : GaloisConnection l u) {s : Set α} (hne : s.Nonempty) (hbdd : BddAbove s) : + l (sSup s) = sSup (l '' s) := by rw [gc.l_csSup hne hbdd, sSup_image'] + +theorem l_ciSup (gc : GaloisConnection l u) {f : ι → α} (hf : BddAbove (range f)) : + l (⨆ i, f i) = ⨆ i, l (f i) := by rw [iSup, gc.l_csSup (range_nonempty _) hf, iSup_range'] + +theorem l_ciSup_set (gc : GaloisConnection l u) {s : Set γ} {f : γ → α} (hf : BddAbove (f '' s)) + (hne : s.Nonempty) : l (⨆ i : s, f i) = ⨆ i : s, l (f i) := by + haveI := hne.to_subtype + rw [image_eq_range] at hf + exact gc.l_ciSup hf + +theorem u_csInf (gc : GaloisConnection l u) {s : Set β} (hne : s.Nonempty) (hbdd : BddBelow s) : + u (sInf s) = ⨅ x : s, u x := + gc.dual.l_csSup hne hbdd + +theorem u_csInf' (gc : GaloisConnection l u) {s : Set β} (hne : s.Nonempty) (hbdd : BddBelow s) : + u (sInf s) = sInf (u '' s) := + gc.dual.l_csSup' hne hbdd + +theorem u_ciInf (gc : GaloisConnection l u) {f : ι → β} (hf : BddBelow (range f)) : + u (⨅ i, f i) = ⨅ i, u (f i) := + gc.dual.l_ciSup hf + +theorem u_ciInf_set (gc : GaloisConnection l u) {s : Set γ} {f : γ → β} (hf : BddBelow (f '' s)) + (hne : s.Nonempty) : u (⨅ i : s, f i) = ⨅ i : s, u (f i) := + gc.dual.l_ciSup_set hf hne + +end GaloisConnection + +namespace OrderIso + +variable [ConditionallyCompleteLattice α] [ConditionallyCompleteLattice β] [Nonempty ι] + +theorem map_csSup (e : α ≃o β) {s : Set α} (hne : s.Nonempty) (hbdd : BddAbove s) : + e (sSup s) = ⨆ x : s, e x := + e.to_galoisConnection.l_csSup hne hbdd + +theorem map_csSup' (e : α ≃o β) {s : Set α} (hne : s.Nonempty) (hbdd : BddAbove s) : + e (sSup s) = sSup (e '' s) := + e.to_galoisConnection.l_csSup' hne hbdd + +theorem map_ciSup (e : α ≃o β) {f : ι → α} (hf : BddAbove (range f)) : + e (⨆ i, f i) = ⨆ i, e (f i) := + e.to_galoisConnection.l_ciSup hf + +theorem map_ciSup_set (e : α ≃o β) {s : Set γ} {f : γ → α} (hf : BddAbove (f '' s)) + (hne : s.Nonempty) : e (⨆ i : s, f i) = ⨆ i : s, e (f i) := + e.to_galoisConnection.l_ciSup_set hf hne + +theorem map_csInf (e : α ≃o β) {s : Set α} (hne : s.Nonempty) (hbdd : BddBelow s) : + e (sInf s) = ⨅ x : s, e x := + e.dual.map_csSup hne hbdd + +theorem map_csInf' (e : α ≃o β) {s : Set α} (hne : s.Nonempty) (hbdd : BddBelow s) : + e (sInf s) = sInf (e '' s) := + e.dual.map_csSup' hne hbdd + +theorem map_ciInf (e : α ≃o β) {f : ι → α} (hf : BddBelow (range f)) : + e (⨅ i, f i) = ⨅ i, e (f i) := + e.dual.map_ciSup hf + +theorem map_ciInf_set (e : α ≃o β) {s : Set γ} {f : γ → α} (hf : BddBelow (f '' s)) + (hne : s.Nonempty) : e (⨅ i : s, f i) = ⨅ i : s, e (f i) := + e.dual.map_ciSup_set hf hne + +end OrderIso + +section WithTopBot + +namespace WithTop +variable [ConditionallyCompleteLinearOrderBot α] {f : ι → α} + +lemma iSup_coe_eq_top : ⨆ x, (f x : WithTop α) = ⊤ ↔ ¬BddAbove (range f) := by + rw [iSup_eq_top, not_bddAbove_iff] + refine ⟨fun hf r => ?_, fun hf a ha => ?_⟩ + · rcases hf r (WithTop.coe_lt_top r) with ⟨i, hi⟩ + exact ⟨f i, ⟨i, rfl⟩, WithTop.coe_lt_coe.mp hi⟩ + · rcases hf (a.untop ha.ne) with ⟨-, ⟨i, rfl⟩, hi⟩ + exact ⟨i, by simpa only [WithTop.coe_untop _ ha.ne] using WithTop.coe_lt_coe.mpr hi⟩ + +lemma iSup_coe_lt_top : ⨆ x, (f x : WithTop α) < ⊤ ↔ BddAbove (range f) := + lt_top_iff_ne_top.trans iSup_coe_eq_top.not_left + +lemma iInf_coe_eq_top : ⨅ x, (f x : WithTop α) = ⊤ ↔ IsEmpty ι := by simp [isEmpty_iff] + +lemma iInf_coe_lt_top : ⨅ i, (f i : WithTop α) < ⊤ ↔ Nonempty ι := by + rw [lt_top_iff_ne_top, Ne, iInf_coe_eq_top, not_isEmpty_iff] + +end WithTop + +end WithTopBot diff --git a/Mathlib/Order/Filter/AtTopBot.lean b/Mathlib/Order/Filter/AtTopBot.lean index 131acf974c0ac..fc058972b73f4 100644 --- a/Mathlib/Order/Filter/AtTopBot.lean +++ b/Mathlib/Order/Filter/AtTopBot.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Jeremy Avigad, Yury Kudryashov, Patrick Massot -/ import Mathlib.Data.Finset.Preimage -import Mathlib.Order.ConditionallyCompleteLattice.Basic +import Mathlib.Order.ConditionallyCompleteLattice.Indexed import Mathlib.Order.Filter.Bases import Mathlib.Order.Filter.Prod import Mathlib.Order.Interval.Set.Disjoint diff --git a/Mathlib/Order/Filter/Extr.lean b/Mathlib/Order/Filter/Extr.lean index e4dbcb43983c5..ba1e7c93d0896 100644 --- a/Mathlib/Order/Filter/Extr.lean +++ b/Mathlib/Order/Filter/Extr.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import Mathlib.Order.Filter.Tendsto -import Mathlib.Order.ConditionallyCompleteLattice.Basic +import Mathlib.Order.ConditionallyCompleteLattice.Indexed import Mathlib.Algebra.Order.Group.Defs /-! diff --git a/Mathlib/Order/PartialSups.lean b/Mathlib/Order/PartialSups.lean index 5fbe0f73945d5..60a016f5f0e41 100644 --- a/Mathlib/Order/PartialSups.lean +++ b/Mathlib/Order/PartialSups.lean @@ -6,7 +6,7 @@ Authors: Kim Morrison import Mathlib.Data.Finset.Lattice import Mathlib.Order.Hom.Basic import Mathlib.Data.Set.Finite -import Mathlib.Order.ConditionallyCompleteLattice.Basic +import Mathlib.Order.ConditionallyCompleteLattice.Indexed /-! # The monotone sequence of partial supremums of a sequence diff --git a/Mathlib/Order/SemiconjSup.lean b/Mathlib/Order/SemiconjSup.lean index ebadc1795ef66..a5fa1eb1ef1fa 100644 --- a/Mathlib/Order/SemiconjSup.lean +++ b/Mathlib/Order/SemiconjSup.lean @@ -6,7 +6,6 @@ Authors: Yury Kudryashov import Mathlib.Algebra.Group.Units.Equiv import Mathlib.Logic.Function.Conjugate import Mathlib.Order.Bounds.OrderIso -import Mathlib.Order.ConditionallyCompleteLattice.Basic import Mathlib.Order.OrdContinuous import Mathlib.Order.RelIso.Group diff --git a/Mathlib/Order/SuccPred/CompleteLinearOrder.lean b/Mathlib/Order/SuccPred/CompleteLinearOrder.lean index 6b8b1de8ce9f3..c8ab06d7b669a 100644 --- a/Mathlib/Order/SuccPred/CompleteLinearOrder.lean +++ b/Mathlib/Order/SuccPred/CompleteLinearOrder.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang -/ import Mathlib.Order.SuccPred.Limit -import Mathlib.Order.ConditionallyCompleteLattice.Basic /-! diff --git a/Mathlib/SetTheory/Cardinal/Basic.lean b/Mathlib/SetTheory/Cardinal/Basic.lean index 90ae6898a01ca..f9e31b0db1b7e 100644 --- a/Mathlib/SetTheory/Cardinal/Basic.lean +++ b/Mathlib/SetTheory/Cardinal/Basic.lean @@ -3,14 +3,15 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro, Floris van Doorn -/ +import Mathlib.Algebra.Order.Ring.Nat import Mathlib.Data.Fintype.BigOperators +import Mathlib.Data.Nat.Cast.Order.Basic import Mathlib.Data.Set.Countable import Mathlib.Logic.Small.Set +import Mathlib.Order.ConditionallyCompleteLattice.Indexed import Mathlib.Order.InitialSeg import Mathlib.Order.SuccPred.CompleteLinearOrder import Mathlib.SetTheory.Cardinal.SchroederBernstein -import Mathlib.Algebra.Order.Ring.Nat -import Mathlib.Data.Nat.Cast.Order.Basic /-! # Cardinal Numbers From 8030444546741c4c639295556201b48730557a08 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Tue, 22 Oct 2024 14:55:24 +0000 Subject: [PATCH 280/505] chore(Analysis): delete "`simp` can prove this" porting notes (#18071) This PR checks all porting notes of the form "`simp` can prove this" and if this is indeed the case (and the proof appears to make sense), removes the associated porting note. --- .../Analysis/BoxIntegral/Partition/Additive.lean | 1 - Mathlib/Analysis/Convex/Join.lean | 1 - Mathlib/Analysis/Fourier/AddCircle.lean | 11 +++++------ Mathlib/Analysis/InnerProductSpace/TwoDim.lean | 2 -- Mathlib/Analysis/Normed/Affine/Isometry.lean | 1 - Mathlib/Analysis/Normed/Field/Basic.lean | 1 - Mathlib/Analysis/Normed/Group/Basic.lean | 8 ++++---- Mathlib/Analysis/Normed/Group/Uniform.lean | 6 ++---- Mathlib/Analysis/Normed/Lp/lpSpace.lean | 1 - .../Analysis/Normed/Operator/LinearIsometry.lean | 13 ------------- Mathlib/Analysis/NormedSpace/Multilinear/Curry.lean | 2 -- Mathlib/Analysis/SpecialFunctions/Exp.lean | 2 -- Mathlib/Analysis/SpecialFunctions/Integrals.lean | 2 -- .../SpecialFunctions/Trigonometric/Angle.lean | 5 ----- .../SpecialFunctions/Trigonometric/Basic.lean | 4 ---- 15 files changed, 11 insertions(+), 49 deletions(-) diff --git a/Mathlib/Analysis/BoxIntegral/Partition/Additive.lean b/Mathlib/Analysis/BoxIntegral/Partition/Additive.lean index 648e9dfc503a9..f82d7bce340fe 100644 --- a/Mathlib/Analysis/BoxIntegral/Partition/Additive.lean +++ b/Mathlib/Analysis/BoxIntegral/Partition/Additive.lean @@ -72,7 +72,6 @@ theorem coe_mk (f h) : ⇑(mk f h : ι →ᵇᵃ[I₀] M) = f := rfl theorem coe_injective : Injective fun (f : ι →ᵇᵃ[I₀] M) x => f x := DFunLike.coe_injective --- Porting note (#10618): was @[simp], now can be proved by `simp` theorem coe_inj {f g : ι →ᵇᵃ[I₀] M} : (f : Box ι → M) = g ↔ f = g := DFunLike.coe_fn_eq theorem sum_partition_boxes (f : ι →ᵇᵃ[I₀] M) (hI : ↑I ≤ I₀) {π : Prepartition I} diff --git a/Mathlib/Analysis/Convex/Join.lean b/Mathlib/Analysis/Convex/Join.lean index 590ebaa4aa614..c2dba9ce9162e 100644 --- a/Mathlib/Analysis/Convex/Join.lean +++ b/Mathlib/Analysis/Convex/Join.lean @@ -59,7 +59,6 @@ theorem convexJoin_singleton_left (t : Set E) (x : E) : theorem convexJoin_singleton_right (s : Set E) (y : E) : convexJoin 𝕜 s {y} = ⋃ x ∈ s, segment 𝕜 x y := by simp [convexJoin] --- Porting note (#10618): simp can prove it theorem convexJoin_singletons (x : E) : convexJoin 𝕜 {x} {y} = segment 𝕜 x y := by simp @[simp] diff --git a/Mathlib/Analysis/Fourier/AddCircle.lean b/Mathlib/Analysis/Fourier/AddCircle.lean index c874a54720ae8..6a5971c612314 100644 --- a/Mathlib/Analysis/Fourier/AddCircle.lean +++ b/Mathlib/Analysis/Fourier/AddCircle.lean @@ -107,7 +107,7 @@ def fourier (n : ℤ) : C(AddCircle T, ℂ) where theorem fourier_apply {n : ℤ} {x : AddCircle T} : fourier n x = toCircle (n • x :) := rfl --- @[simp] -- Porting note: simp normal form is `fourier_coe_apply'` +-- simp normal form is `fourier_coe_apply'` theorem fourier_coe_apply {n : ℤ} {x : ℝ} : fourier n (x : AddCircle T) = Complex.exp (2 * π * Complex.I * n * x / T) := by rw [fourier_apply, ← QuotientAddGroup.mk_zsmul, toCircle, Function.Periodic.lift_coe, @@ -121,7 +121,7 @@ theorem fourier_coe_apply' {n : ℤ} {x : ℝ} : toCircle (n • (x : AddCircle T) :) = Complex.exp (2 * π * Complex.I * n * x / T) := by rw [← fourier_apply]; exact fourier_coe_apply --- @[simp] -- Porting note: simp normal form is `fourier_zero'` +-- simp normal form is `fourier_zero'` theorem fourier_zero {x : AddCircle T} : fourier 0 x = 1 := by induction x using QuotientAddGroup.induction_on simp only [fourier_coe_apply] @@ -132,15 +132,14 @@ theorem fourier_zero' {x : AddCircle T} : @toCircle T 0 = (1 : ℂ) := by have : fourier 0 x = @toCircle T 0 := by rw [fourier_apply, zero_smul] rw [← this]; exact fourier_zero --- @[simp] -- Porting note: simp normal form is *also* `fourier_zero'` +-- simp normal form is *also* `fourier_zero'` theorem fourier_eval_zero (n : ℤ) : fourier n (0 : AddCircle T) = 1 := by rw [← QuotientAddGroup.mk_zero, fourier_coe_apply, Complex.ofReal_zero, mul_zero, zero_div, Complex.exp_zero] --- @[simp] -- Porting note (#10618): simp can prove this theorem fourier_one {x : AddCircle T} : fourier 1 x = toCircle x := by rw [fourier_apply, one_zsmul] --- @[simp] -- Porting note: simp normal form is `fourier_neg'` +-- simp normal form is `fourier_neg'` theorem fourier_neg {n : ℤ} {x : AddCircle T} : fourier (-n) x = conj (fourier n x) := by induction x using QuotientAddGroup.induction_on simp_rw [fourier_apply, toCircle] @@ -152,7 +151,7 @@ theorem fourier_neg {n : ℤ} {x : AddCircle T} : fourier (-n) x = conj (fourier theorem fourier_neg' {n : ℤ} {x : AddCircle T} : @toCircle T (-(n • x)) = conj (fourier n x) := by rw [← neg_smul, ← fourier_apply]; exact fourier_neg --- @[simp] -- Porting note: simp normal form is `fourier_add'` +-- simp normal form is `fourier_add'` theorem fourier_add {m n : ℤ} {x : AddCircle T} : fourier (m+n) x = fourier m x * fourier n x := by simp_rw [fourier_apply, add_zsmul, toCircle_add, Circle.coe_mul] diff --git a/Mathlib/Analysis/InnerProductSpace/TwoDim.lean b/Mathlib/Analysis/InnerProductSpace/TwoDim.lean index 329dd441e2adc..a68bcb961c62e 100644 --- a/Mathlib/Analysis/InnerProductSpace/TwoDim.lean +++ b/Mathlib/Analysis/InnerProductSpace/TwoDim.lean @@ -260,7 +260,6 @@ theorem rightAngleRotation_symm : rw [rightAngleRotation] exact LinearIsometryEquiv.toLinearIsometry_injective rfl --- @[simp] -- Porting note (#10618): simp already proves this theorem inner_rightAngleRotation_self (x : E) : ⟪J x, x⟫ = 0 := by simp theorem inner_rightAngleRotation_swap (x y : E) : ⟪x, J y⟫ = -⟪J x, y⟫ := by simp @@ -279,7 +278,6 @@ theorem areaForm_rightAngleRotation_left (x y : E) : ω (J x) y = -⟪x, y⟫ := theorem areaForm_rightAngleRotation_right (x y : E) : ω x (J y) = ⟪x, y⟫ := by rw [← o.inner_rightAngleRotation_left, o.inner_comp_rightAngleRotation] --- @[simp] -- Porting note (#10618): simp already proves this theorem areaForm_comp_rightAngleRotation (x y : E) : ω (J x) (J y) = ω x y := by simp @[simp] diff --git a/Mathlib/Analysis/Normed/Affine/Isometry.lean b/Mathlib/Analysis/Normed/Affine/Isometry.lean index 67f6a72b44673..d6db878a74fdd 100644 --- a/Mathlib/Analysis/Normed/Affine/Isometry.lean +++ b/Mathlib/Analysis/Normed/Affine/Isometry.lean @@ -541,7 +541,6 @@ protected theorem injective : Injective e := protected theorem surjective : Surjective e := e.1.surjective --- @[simp] Porting note (#10618): simp can prove this theorem map_eq_iff {x y : P} : e x = e y ↔ x = y := e.injective.eq_iff diff --git a/Mathlib/Analysis/Normed/Field/Basic.lean b/Mathlib/Analysis/Normed/Field/Basic.lean index d548024cd56e3..da137b5e2cc71 100644 --- a/Mathlib/Analysis/Normed/Field/Basic.lean +++ b/Mathlib/Analysis/Normed/Field/Basic.lean @@ -922,7 +922,6 @@ namespace NNReal open NNReal --- porting note (#10618): removed `@[simp]` because `simp` can prove this theorem norm_eq (x : ℝ≥0) : ‖(x : ℝ)‖ = x := by rw [Real.norm_eq_abs, x.abs_eq] @[simp] diff --git a/Mathlib/Analysis/Normed/Group/Basic.lean b/Mathlib/Analysis/Normed/Group/Basic.lean index 31e3b6bff3fe5..50718d0713f3c 100644 --- a/Mathlib/Analysis/Normed/Group/Basic.lean +++ b/Mathlib/Analysis/Normed/Group/Basic.lean @@ -516,14 +516,14 @@ theorem mem_ball_iff_norm'' : b ∈ ball a r ↔ ‖b / a‖ < r := by rw [mem_b @[to_additive mem_ball_iff_norm'] theorem mem_ball_iff_norm''' : b ∈ ball a r ↔ ‖a / b‖ < r := by rw [mem_ball', dist_eq_norm_div] -@[to_additive] -- Porting note (#10618): `simp` can prove it +@[to_additive] theorem mem_ball_one_iff : a ∈ ball (1 : E) r ↔ ‖a‖ < r := by rw [mem_ball, dist_one_right] @[to_additive mem_closedBall_iff_norm] theorem mem_closedBall_iff_norm'' : b ∈ closedBall a r ↔ ‖b / a‖ ≤ r := by rw [mem_closedBall, dist_eq_norm_div] -@[to_additive] -- Porting note (#10618): `simp` can prove it +@[to_additive] theorem mem_closedBall_one_iff : a ∈ closedBall (1 : E) r ↔ ‖a‖ ≤ r := by rw [mem_closedBall, dist_one_right] @@ -1073,11 +1073,11 @@ theorem pow_mem_ball {n : ℕ} (hn : 0 < n) (h : a ∈ ball b r) : a ^ n ∈ bal rw [nsmul_eq_mul] nlinarith -@[to_additive] -- Porting note (#10618): `simp` can prove this +@[to_additive] theorem mul_mem_closedBall_mul_iff {c : E} : a * c ∈ closedBall (b * c) r ↔ a ∈ closedBall b r := by simp only [mem_closedBall, dist_eq_norm_div, mul_div_mul_right_eq_div] -@[to_additive] -- Porting note (#10618): `simp` can prove this +@[to_additive] theorem mul_mem_ball_mul_iff {c : E} : a * c ∈ ball (b * c) r ↔ a ∈ ball b r := by simp only [mem_ball, dist_eq_norm_div, mul_div_mul_right_eq_div] diff --git a/Mathlib/Analysis/Normed/Group/Uniform.lean b/Mathlib/Analysis/Normed/Group/Uniform.lean index f8e72a4535b45..3f63d8ab90c45 100644 --- a/Mathlib/Analysis/Normed/Group/Uniform.lean +++ b/Mathlib/Analysis/Normed/Group/Uniform.lean @@ -191,13 +191,11 @@ theorem dist_self_mul_right (a b : E) : dist a (a * b) = ‖b‖ := by theorem dist_self_mul_left (a b : E) : dist (a * b) a = ‖b‖ := by rw [dist_comm, dist_self_mul_right] -@[to_additive (attr := simp 1001)] --- porting note (#10618): increase priority because `simp` can prove this +@[to_additive (attr := simp 1001)] -- Increase priority because `simp` can prove this theorem dist_self_div_right (a b : E) : dist a (a / b) = ‖b‖ := by rw [div_eq_mul_inv, dist_self_mul_right, norm_inv'] -@[to_additive (attr := simp 1001)] --- porting note (#10618): increase priority because `simp` can prove this +@[to_additive (attr := simp 1001)] -- Increase priority because `simp` can prove this theorem dist_self_div_left (a b : E) : dist (a / b) a = ‖b‖ := by rw [dist_comm, dist_self_div_right] diff --git a/Mathlib/Analysis/Normed/Lp/lpSpace.lean b/Mathlib/Analysis/Normed/Lp/lpSpace.lean index 0c44871b94833..a3ebaed3c2b34 100644 --- a/Mathlib/Analysis/Normed/Lp/lpSpace.lean +++ b/Mathlib/Analysis/Normed/Lp/lpSpace.lean @@ -336,7 +336,6 @@ theorem coeFn_neg (f : lp E p) : ⇑(-f) = -f := theorem coeFn_add (f g : lp E p) : ⇑(f + g) = f + g := rfl --- porting note (#10618): removed `@[simp]` because `simp` can prove this theorem coeFn_sum {ι : Type*} (f : ι → lp E p) (s : Finset ι) : ⇑(∑ i ∈ s, f i) = ∑ i ∈ s, ⇑(f i) := by simp diff --git a/Mathlib/Analysis/Normed/Operator/LinearIsometry.lean b/Mathlib/Analysis/Normed/Operator/LinearIsometry.lean index 0146ea7262fd7..facb95753719a 100644 --- a/Mathlib/Analysis/Normed/Operator/LinearIsometry.lean +++ b/Mathlib/Analysis/Normed/Operator/LinearIsometry.lean @@ -171,27 +171,21 @@ theorem ext {f g : E →ₛₗᵢ[σ₁₂] E₂} (h : ∀ x, f x = g x) : f = g variable [FunLike 𝓕 E E₂] --- @[simp] -- Porting note (#10618): simp can prove this protected theorem map_zero : f 0 = 0 := f.toLinearMap.map_zero --- @[simp] -- Porting note (#10618): simp can prove this protected theorem map_add (x y : E) : f (x + y) = f x + f y := f.toLinearMap.map_add x y --- @[simp] -- Porting note (#10618): simp can prove this protected theorem map_neg (x : E) : f (-x) = -f x := f.toLinearMap.map_neg x --- @[simp] -- Porting note (#10618): simp can prove this protected theorem map_sub (x y : E) : f (x - y) = f x - f y := f.toLinearMap.map_sub x y --- @[simp] -- Porting note (#10618): simp can prove this protected theorem map_smulₛₗ (c : R) (x : E) : f (c • x) = σ₁₂ c • f x := f.toLinearMap.map_smulₛₗ c x --- @[simp] -- Porting note (#10618): simp can prove this protected theorem map_smul [Module R E₂] (f : E →ₗᵢ[R] E₂) (c : R) (x : E) : f (c • x) = c • f x := f.toLinearMap.map_smul c x @@ -641,7 +635,6 @@ theorem apply_symm_apply (x : E₂) : e (e.symm x) = x := theorem symm_apply_apply (x : E) : e.symm (e x) = x := e.toLinearEquiv.symm_apply_apply x --- @[simp] -- Porting note (#10618): simp can prove this theorem map_eq_zero_iff {x : E} : e x = 0 ↔ x = 0 := e.toLinearEquiv.map_eq_zero_iff @@ -805,23 +798,18 @@ theorem coe_coe : ⇑(e : E ≃SL[σ₁₂] E₂) = e := theorem coe_coe'' : ⇑(e : E →SL[σ₁₂] E₂) = e := rfl --- @[simp] -- Porting note (#10618): simp can prove this theorem map_zero : e 0 = 0 := e.1.map_zero --- @[simp] -- Porting note (#10618): simp can prove this theorem map_add (x y : E) : e (x + y) = e x + e y := e.1.map_add x y --- @[simp] -- Porting note (#10618): simp can prove this theorem map_sub (x y : E) : e (x - y) = e x - e y := e.1.map_sub x y --- @[simp] -- Porting note (#10618): simp can prove this theorem map_smulₛₗ (c : R) (x : E) : e (c • x) = σ₁₂ c • e x := e.1.map_smulₛₗ c x --- @[simp] -- Porting note (#10618): simp can prove this theorem map_smul [Module R E₂] {e : E ≃ₗᵢ[R] E₂} (c : R) (x : E) : e (c • x) = c • e x := e.1.map_smul c x @@ -846,7 +834,6 @@ protected theorem injective : Injective e := protected theorem surjective : Surjective e := e.1.surjective --- @[simp] -- Porting note (#10618): simp can prove this theorem map_eq_iff {x y : E} : e x = e y ↔ x = y := e.injective.eq_iff diff --git a/Mathlib/Analysis/NormedSpace/Multilinear/Curry.lean b/Mathlib/Analysis/NormedSpace/Multilinear/Curry.lean index 702bd8dea02d9..2ce89faee16da 100644 --- a/Mathlib/Analysis/NormedSpace/Multilinear/Curry.lean +++ b/Mathlib/Analysis/NormedSpace/Multilinear/Curry.lean @@ -580,7 +580,6 @@ theorem curryFinFinset_symm_apply (hk : s.card = k) (hl : sᶜ.card = l) m <| finSumEquivOfFinset hk hl (Sum.inr i) := rfl --- @[simp] -- Porting note (#10618): simp removed: simp can reduce LHS theorem curryFinFinset_symm_apply_piecewise_const (hk : s.card = k) (hl : sᶜ.card = l) (f : G[×k]→L[𝕜] G[×l]→L[𝕜] G') (x y : G) : (curryFinFinset 𝕜 G G' hk hl).symm f (s.piecewise (fun _ => x) fun _ => y) = @@ -593,7 +592,6 @@ theorem curryFinFinset_symm_apply_const (hk : s.card = k) (hl : sᶜ.card = l) ((curryFinFinset 𝕜 G G' hk hl).symm f fun _ => x) = f (fun _ => x) fun _ => x := rfl --- @[simp] -- Porting note (#10618): simp removed: simp can reduce LHS theorem curryFinFinset_apply_const (hk : s.card = k) (hl : sᶜ.card = l) (f : G[×n]→L[𝕜] G') (x y : G) : (curryFinFinset 𝕜 G G' hk hl f (fun _ => x) fun _ => y) = f (s.piecewise (fun _ => x) fun _ => y) := by diff --git a/Mathlib/Analysis/SpecialFunctions/Exp.lean b/Mathlib/Analysis/SpecialFunctions/Exp.lean index 9a02b4d3dc803..bd9a8f4a616de 100644 --- a/Mathlib/Analysis/SpecialFunctions/Exp.lean +++ b/Mathlib/Analysis/SpecialFunctions/Exp.lean @@ -406,8 +406,6 @@ theorem isLittleO_exp_comp_exp_comp {f g : α → ℝ} : simp only [isLittleO_iff_tendsto, exp_ne_zero, ← exp_sub, ← tendsto_neg_atTop_iff, false_imp_iff, imp_true_iff, tendsto_exp_comp_nhds_zero, neg_sub] --- Porting note (#10618): @[simp] can prove: by simp only [@Asymptotics.isLittleO_one_left_iff, --- Real.norm_eq_abs, Real.abs_exp, @Real.tendsto_exp_comp_atTop] theorem isLittleO_one_exp_comp {f : α → ℝ} : ((fun _ => 1 : α → ℝ) =o[l] fun x => exp (f x)) ↔ Tendsto f l atTop := by simp only [← exp_zero, isLittleO_exp_comp_exp_comp, sub_zero] diff --git a/Mathlib/Analysis/SpecialFunctions/Integrals.lean b/Mathlib/Analysis/SpecialFunctions/Integrals.lean index 31f51791cdb31..e11cef64dcdb8 100644 --- a/Mathlib/Analysis/SpecialFunctions/Integrals.lean +++ b/Mathlib/Analysis/SpecialFunctions/Integrals.lean @@ -206,7 +206,6 @@ theorem integrableOn_Ioo_cpow_iff {s : ℂ} {t : ℝ} (ht : 0 < t) : theorem intervalIntegrable_id : IntervalIntegrable (fun x => x) μ a b := continuous_id.intervalIntegrable a b --- @[simp] -- Porting note (#10618): simp can prove this theorem intervalIntegrable_const : IntervalIntegrable (fun _ => c) μ a b := continuous_const.intervalIntegrable a b @@ -409,7 +408,6 @@ theorem integral_id : ∫ x in a..b, x = (b ^ 2 - a ^ 2) / 2 := by norm_num at this exact this --- @[simp] -- Porting note (#10618): simp can prove this theorem integral_one : (∫ _ in a..b, (1 : ℝ)) = b - a := by simp only [mul_one, smul_eq_mul, integral_const] diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Angle.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Angle.lean index c5ed691712471..6b91d16d57e6c 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Angle.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Angle.lean @@ -127,11 +127,9 @@ theorem two_nsmul_coe_div_two (θ : ℝ) : (2 : ℕ) • (↑(θ / 2) : Angle) = theorem two_zsmul_coe_div_two (θ : ℝ) : (2 : ℤ) • (↑(θ / 2) : Angle) = θ := by rw [← coe_zsmul, two_zsmul, add_halves] --- Porting note (#10618): @[simp] can prove it theorem two_nsmul_neg_pi_div_two : (2 : ℕ) • (↑(-π / 2) : Angle) = π := by rw [two_nsmul_coe_div_two, coe_neg, neg_coe_pi] --- Porting note (#10618): @[simp] can prove it theorem two_zsmul_neg_pi_div_two : (2 : ℤ) • (↑(-π / 2) : Angle) = π := by rw [two_zsmul, ← two_nsmul, two_nsmul_neg_pi_div_two] @@ -303,7 +301,6 @@ theorem sin_eq_iff_eq_or_add_eq_pi {θ ψ : Angle} : sin θ = sin ψ ↔ θ = ψ @[simp] theorem sin_zero : sin (0 : Angle) = 0 := by rw [← coe_zero, sin_coe, Real.sin_zero] --- Porting note (#10618): @[simp] can prove it theorem sin_coe_pi : sin (π : Angle) = 0 := by rw [sin_coe, Real.sin_pi] theorem sin_eq_zero_iff {θ : Angle} : sin θ = 0 ↔ θ = 0 ∨ θ = π := by @@ -335,7 +332,6 @@ theorem sin_sub_pi (θ : Angle) : sin (θ - π) = -sin θ := @[simp] theorem cos_zero : cos (0 : Angle) = 1 := by rw [← coe_zero, cos_coe, Real.cos_zero] --- Porting note (#10618): @[simp] can prove it theorem cos_coe_pi : cos (π : Angle) = -1 := by rw [cos_coe, Real.cos_pi] @[simp] @@ -658,7 +654,6 @@ theorem tan_coe (x : ℝ) : tan (x : Angle) = Real.tan x := by @[simp] theorem tan_zero : tan (0 : Angle) = 0 := by rw [← coe_zero, tan_coe, Real.tan_zero] --- Porting note (#10618): @[simp] can now prove it theorem tan_coe_pi : tan (π : Angle) = 0 := by rw [tan_coe, Real.tan_pi] theorem tan_periodic : Function.Periodic tan (π : Angle) := by diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean index b09fb6620f62e..4c002f27c027e 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean @@ -380,19 +380,15 @@ theorem cos_int_mul_pi_sub (x : ℝ) (n : ℤ) : cos (n * π - x) = (-1) ^ n * c theorem cos_nat_mul_pi_sub (x : ℝ) (n : ℕ) : cos (n * π - x) = (-1) ^ n * cos x := cos_neg x ▸ cos_antiperiodic.nat_mul_sub_eq n --- Porting note (#10618): was @[simp], but simp can prove it theorem cos_nat_mul_two_pi_add_pi (n : ℕ) : cos (n * (2 * π) + π) = -1 := by simpa only [cos_zero] using (cos_periodic.nat_mul n).add_antiperiod_eq cos_antiperiodic --- Porting note (#10618): was @[simp], but simp can prove it theorem cos_int_mul_two_pi_add_pi (n : ℤ) : cos (n * (2 * π) + π) = -1 := by simpa only [cos_zero] using (cos_periodic.int_mul n).add_antiperiod_eq cos_antiperiodic --- Porting note (#10618): was @[simp], but simp can prove it theorem cos_nat_mul_two_pi_sub_pi (n : ℕ) : cos (n * (2 * π) - π) = -1 := by simpa only [cos_zero] using (cos_periodic.nat_mul n).sub_antiperiod_eq cos_antiperiodic --- Porting note (#10618): was @[simp], but simp can prove it theorem cos_int_mul_two_pi_sub_pi (n : ℤ) : cos (n * (2 * π) - π) = -1 := by simpa only [cos_zero] using (cos_periodic.int_mul n).sub_antiperiod_eq cos_antiperiodic From 9d61276a94857b7c19f2e62610a8939880226b02 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Tue, 22 Oct 2024 14:55:26 +0000 Subject: [PATCH 281/505] chore(CategoryTheory): delete "`simp` can prove this" porting notes (#18072) This PR checks all porting notes of the form "`simp` can prove this" and if this is indeed the case (and the proof appears to make sense), removes the associated porting note. --- Mathlib/CategoryTheory/Abelian/Images.lean | 1 - Mathlib/CategoryTheory/Abelian/NonPreadditive.lean | 1 - Mathlib/CategoryTheory/Adjunction/Unique.lean | 2 -- Mathlib/CategoryTheory/Comma/Arrow.lean | 2 -- Mathlib/CategoryTheory/HomCongr.lean | 2 -- Mathlib/CategoryTheory/Limits/Cones.lean | 2 +- Mathlib/CategoryTheory/Limits/Opposites.lean | 8 -------- Mathlib/CategoryTheory/Limits/Shapes/BinaryProducts.lean | 1 - Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean | 2 +- Mathlib/CategoryTheory/Limits/Shapes/Equalizers.lean | 3 --- Mathlib/CategoryTheory/Limits/Shapes/Kernels.lean | 2 -- Mathlib/CategoryTheory/Limits/Shapes/Multiequalizer.lean | 4 ++-- .../Limits/Shapes/Pullback/HasPullback.lean | 6 ------ Mathlib/CategoryTheory/Limits/Shapes/Pullback/Iso.lean | 4 ---- Mathlib/CategoryTheory/Limits/Shapes/WidePullbacks.lean | 3 --- Mathlib/CategoryTheory/PUnit.lean | 2 -- .../CategoryTheory/Preadditive/InjectiveResolution.lean | 2 -- .../CategoryTheory/Preadditive/ProjectiveResolution.lean | 1 - 18 files changed, 4 insertions(+), 44 deletions(-) diff --git a/Mathlib/CategoryTheory/Abelian/Images.lean b/Mathlib/CategoryTheory/Abelian/Images.lean index c753cc0c81887..1bed70213ca90 100644 --- a/Mathlib/CategoryTheory/Abelian/Images.lean +++ b/Mathlib/CategoryTheory/Abelian/Images.lean @@ -49,7 +49,6 @@ protected abbrev image.ι : Abelian.image f ⟶ Q := protected abbrev factorThruImage : P ⟶ Abelian.image f := kernel.lift (cokernel.π f) f <| cokernel.condition f --- Porting note (#10618): simp can prove this and reassoc version, removed tags /-- `f` factors through its image via the canonical morphism `p`. -/ protected theorem image.fac : Abelian.factorThruImage f ≫ image.ι f = f := kernel.lift_ι _ _ _ diff --git a/Mathlib/CategoryTheory/Abelian/NonPreadditive.lean b/Mathlib/CategoryTheory/Abelian/NonPreadditive.lean index 810dadb3c8117..8f4d7dda4a14c 100644 --- a/Mathlib/CategoryTheory/Abelian/NonPreadditive.lean +++ b/Mathlib/CategoryTheory/Abelian/NonPreadditive.lean @@ -265,7 +265,6 @@ abbrev σ {A : C} : A ⨯ A ⟶ A := end --- Porting note (#10618): simp can prove these @[reassoc] theorem diag_σ {X : C} : diag X ≫ σ = 0 := by rw [cokernel.condition_assoc, zero_comp] diff --git a/Mathlib/CategoryTheory/Adjunction/Unique.lean b/Mathlib/CategoryTheory/Adjunction/Unique.lean index 7b1a98e824fa1..ec1dcbe176388 100644 --- a/Mathlib/CategoryTheory/Adjunction/Unique.lean +++ b/Mathlib/CategoryTheory/Adjunction/Unique.lean @@ -32,7 +32,6 @@ attribute [local simp] homEquiv_unit homEquiv_counit def leftAdjointUniq {F F' : C ⥤ D} {G : D ⥤ C} (adj1 : F ⊣ G) (adj2 : F' ⊣ G) : F ≅ F' := ((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) (x : C) : adj1.homEquiv _ _ ((leftAdjointUniq adj1 adj2).hom.app x) = adj2.unit.app x := by simp [leftAdjointUniq] @@ -96,7 +95,6 @@ theorem leftAdjointUniq_refl {F : C ⥤ D} {G : D ⥤ C} (adj1 : F ⊣ G) : def rightAdjointUniq {F : C ⥤ D} {G G' : D ⥤ C} (adj1 : F ⊣ G) (adj2 : F ⊣ G') : G ≅ G' := 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) (adj2 : F ⊣ G') (x : D) : (adj2.homEquiv _ _).symm ((rightAdjointUniq adj1 adj2).hom.app x) = adj1.counit.app x := by diff --git a/Mathlib/CategoryTheory/Comma/Arrow.lean b/Mathlib/CategoryTheory/Comma/Arrow.lean index 9ab800f7f7cc4..2d8f03930b499 100644 --- a/Mathlib/CategoryTheory/Comma/Arrow.lean +++ b/Mathlib/CategoryTheory/Comma/Arrow.lean @@ -184,11 +184,9 @@ theorem inv_left [IsIso sq] : (inv sq).left = inv sq.left := theorem inv_right [IsIso sq] : (inv sq).right = inv sq.right := IsIso.eq_inv_of_hom_inv_id <| by rw [← Comma.comp_right, IsIso.hom_inv_id, id_right] -/- Porting note (#10618): simp can prove this so removed @[simp] -/ theorem left_hom_inv_right [IsIso sq] : sq.left ≫ g.hom ≫ inv sq.right = f.hom := by simp only [← Category.assoc, IsIso.comp_inv_eq, w] --- simp proves this theorem inv_left_hom_right [IsIso sq] : inv sq.left ≫ f.hom ≫ sq.right = g.hom := by simp only [w, IsIso.inv_comp_eq] diff --git a/Mathlib/CategoryTheory/HomCongr.lean b/Mathlib/CategoryTheory/HomCongr.lean index ac21e83f53176..f2619de2c30f1 100644 --- a/Mathlib/CategoryTheory/HomCongr.lean +++ b/Mathlib/CategoryTheory/HomCongr.lean @@ -46,10 +46,8 @@ def homCongr {X Y X₁ Y₁ : C} (α : X ≅ X₁) (β : Y ≅ Y₁) : (X ⟶ Y) theorem homCongr_comp {X Y Z X₁ Y₁ Z₁ : C} (α : X ≅ X₁) (β : Y ≅ Y₁) (γ : Z ≅ Z₁) (f : X ⟶ Y) (g : Y ⟶ Z) : α.homCongr γ (f ≫ g) = α.homCongr β f ≫ β.homCongr γ g := by simp -/- Porting note (#10618): removed `@[simp]`; simp can prove this -/ theorem homCongr_refl {X Y : C} (f : X ⟶ Y) : (Iso.refl X).homCongr (Iso.refl Y) f = f := by simp -/- Porting note (#10618): removed `@[simp]`; simp can prove this -/ theorem homCongr_trans {X₁ Y₁ X₂ Y₂ X₃ Y₃ : C} (α₁ : X₁ ≅ X₂) (β₁ : Y₁ ≅ Y₂) (α₂ : X₂ ≅ X₃) (β₂ : Y₂ ≅ Y₃) (f : X₁ ⟶ Y₁) : (α₁ ≪≫ α₂).homCongr (β₁ ≪≫ β₂) f = (α₁.homCongr β₁).trans (α₂.homCongr β₂) f := by simp diff --git a/Mathlib/CategoryTheory/Limits/Cones.lean b/Mathlib/CategoryTheory/Limits/Cones.lean index 36d559c2f7911..6fee46c53e64c 100644 --- a/Mathlib/CategoryTheory/Limits/Cones.lean +++ b/Mathlib/CategoryTheory/Limits/Cones.lean @@ -162,7 +162,7 @@ instance inhabitedCocone (F : Discrete PUnit ⥤ C) : Inhabited (Cocone F) := } }⟩ -@[reassoc] -- @[simp] -- Porting note (#10618): simp can prove this +@[reassoc] theorem Cocone.w {F : J ⥤ C} (c : Cocone F) {j j' : J} (f : j ⟶ j') : F.map f ≫ c.ι.app j' = c.ι.app j := by rw [c.ι.naturality f] diff --git a/Mathlib/CategoryTheory/Limits/Opposites.lean b/Mathlib/CategoryTheory/Limits/Opposites.lean index c630f89557f36..f53eb0808b891 100644 --- a/Mathlib/CategoryTheory/Limits/Opposites.lean +++ b/Mathlib/CategoryTheory/Limits/Opposites.lean @@ -635,11 +635,9 @@ def unop {X Y Z : Cᵒᵖ} {f : X ⟶ Y} {g : X ⟶ Z} (c : PushoutCocone f g) : ((Cocones.precompose (opCospan f.unop g.unop).hom).obj (Cocone.whisker walkingCospanOpEquiv.functor c)) --- Porting note (#10618): removed simp attribute as the equality can already be obtained by simp theorem unop_fst {X Y Z : Cᵒᵖ} {f : X ⟶ Y} {g : X ⟶ Z} (c : PushoutCocone f g) : c.unop.fst = c.inl.unop := by simp --- Porting note (#10618): removed simp attribute as the equality can already be obtained by simp theorem unop_snd {X Y Z : Cᵒᵖ} {f : X ⟶ Y} {g : X ⟶ Z} (c : PushoutCocone f g) : c.unop.snd = c.inr.unop := by aesop_cat @@ -650,11 +648,9 @@ def op {X Y Z : C} {f : X ⟶ Y} {g : X ⟶ Z} (c : PushoutCocone f g) : Pullbac (Cones.postcompose (cospanOp f g).symm.hom).obj (Cone.whisker walkingSpanOpEquiv.inverse (Cocone.op c)) --- Porting note (#10618): removed simp attribute as the equality can already be obtained by simp theorem op_fst {X Y Z : C} {f : X ⟶ Y} {g : X ⟶ Z} (c : PushoutCocone f g) : c.op.fst = c.inl.op := by aesop_cat --- Porting note (#10618): removed simp attribute as the equality can already be obtained by simp theorem op_snd {X Y Z : C} {f : X ⟶ Y} {g : X ⟶ Z} (c : PushoutCocone f g) : c.op.snd = c.inr.op := by aesop_cat @@ -671,11 +667,9 @@ def unop {X Y Z : Cᵒᵖ} {f : X ⟶ Z} {g : Y ⟶ Z} (c : PullbackCone f g) : ((Cones.postcompose (opSpan f.unop g.unop).symm.hom).obj (Cone.whisker walkingSpanOpEquiv.functor c)) --- Porting note (#10618): removed simp attribute as the equality can already be obtained by simp theorem unop_inl {X Y Z : Cᵒᵖ} {f : X ⟶ Z} {g : Y ⟶ Z} (c : PullbackCone f g) : c.unop.inl = c.fst.unop := by aesop_cat --- Porting note (#10618): removed simp attribute as the equality can already be obtained by simp theorem unop_inr {X Y Z : Cᵒᵖ} {f : X ⟶ Z} {g : Y ⟶ Z} (c : PullbackCone f g) : c.unop.inr = c.snd.unop := by aesop_cat @@ -685,11 +679,9 @@ def op {X Y Z : C} {f : X ⟶ Z} {g : Y ⟶ Z} (c : PullbackCone f g) : PushoutC (Cocones.precompose (spanOp f g).hom).obj (Cocone.whisker walkingCospanOpEquiv.inverse (Cone.op c)) --- Porting note (#10618): removed simp attribute as the equality can already be obtained by simp theorem op_inl {X Y Z : C} {f : X ⟶ Z} {g : Y ⟶ Z} (c : PullbackCone f g) : c.op.inl = c.fst.op := by aesop_cat --- Porting note (#10618): removed simp attribute as the equality can already be obtained by simp theorem op_inr {X Y Z : C} {f : X ⟶ Z} {g : Y ⟶ Z} (c : PullbackCone f g) : c.op.inr = c.snd.op := by aesop_cat diff --git a/Mathlib/CategoryTheory/Limits/Shapes/BinaryProducts.lean b/Mathlib/CategoryTheory/Limits/Shapes/BinaryProducts.lean index 0d8529d45ba77..c399e2fe79062 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/BinaryProducts.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/BinaryProducts.lean @@ -553,7 +553,6 @@ theorem prod.lift_fst {W X Y : C} [HasBinaryProduct X Y] (f : W ⟶ X) (g : W prod.lift f g ≫ prod.fst = f := limit.lift_π _ _ --- Porting note (#10618): simp removes as simp can prove this @[reassoc] theorem prod.lift_snd {W X Y : C} [HasBinaryProduct X Y] (f : W ⟶ X) (g : W ⟶ Y) : prod.lift f g ≫ prod.snd = g := diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean b/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean index f290e6be94303..5bf06db5fdc26 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean @@ -755,7 +755,7 @@ theorem biproduct.toSubtype_eq_desc [DecidablePred p] : biproduct.desc fun j => if h : p j then biproduct.ι (Subtype.restrict p f) ⟨j, h⟩ else 0 := biproduct.hom_ext' _ _ (by simp) -@[reassoc] -- Porting note (#10618): simp can prove both versions +@[reassoc] theorem biproduct.ι_toSubtype_subtype (j : Subtype p) : biproduct.ι f j ≫ biproduct.toSubtype f p = biproduct.ι (Subtype.restrict p f) j := by ext diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Equalizers.lean b/Mathlib/CategoryTheory/Limits/Shapes/Equalizers.lean index 0bc2aff0e4a46..36a0a6e065433 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Equalizers.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Equalizers.lean @@ -719,7 +719,6 @@ variable {f g} noncomputable abbrev equalizer.lift {W : C} (k : W ⟶ X) (h : k ≫ f = k ≫ g) : W ⟶ equalizer f g := limit.lift (parallelPair f g) (Fork.ofι k h) --- Porting note (#10618): removed simp since simp can prove this and the reassoc version @[reassoc] theorem equalizer.lift_ι {W : C} (k : W ⟶ X) (h : k ≫ f = k ≫ g) : equalizer.lift k h ≫ equalizer.ι f g = k := @@ -850,7 +849,6 @@ noncomputable abbrev coequalizer.cofork : Cofork f g := theorem coequalizer.cofork_π : (coequalizer.cofork f g).π = coequalizer.π f g := rfl --- Porting note (#10618): simp can prove this, simp removed theorem coequalizer.cofork_ι_app_one : (coequalizer.cofork f g).ι.app one = coequalizer.π f g := rfl @@ -871,7 +869,6 @@ noncomputable abbrev coequalizer.desc {W : C} (k : Y ⟶ W) (h : f ≫ k = g ≫ coequalizer f g ⟶ W := colimit.desc (parallelPair f g) (Cofork.ofπ k h) --- Porting note (#10618): removing simp since simp can prove this and reassoc version @[reassoc] theorem coequalizer.π_desc {W : C} (k : Y ⟶ W) (h : f ≫ k = g ≫ k) : coequalizer.π f g ≫ coequalizer.desc k h = k := diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Kernels.lean b/Mathlib/CategoryTheory/Limits/Shapes/Kernels.lean index 9015e814eea12..f03cd551569a5 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Kernels.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Kernels.lean @@ -81,7 +81,6 @@ variable {f} theorem KernelFork.condition (s : KernelFork f) : Fork.ι s ≫ f = 0 := by rw [Fork.condition, HasZeroMorphisms.comp_zero] --- Porting note (#10618): simp can prove this, removed simp tag theorem KernelFork.app_one (s : KernelFork f) : s.π.app one = 0 := by simp [Fork.app_one_eq_ι_comp_right] @@ -516,7 +515,6 @@ variable {f} theorem CokernelCofork.condition (s : CokernelCofork f) : f ≫ s.π = 0 := by rw [Cofork.condition, zero_comp] --- Porting note (#10618): simp can prove this, removed simp tag theorem CokernelCofork.π_eq_zero (s : CokernelCofork f) : s.ι.app zero = 0 := by simp [Cofork.app_zero_eq_comp_π_right] diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Multiequalizer.lean b/Mathlib/CategoryTheory/Limits/Shapes/Multiequalizer.lean index d880d204adea3..d38347887cbc5 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Multiequalizer.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Multiequalizer.lean @@ -722,7 +722,7 @@ abbrev lift (W : C) (k : ∀ a, W ⟶ I.left a) (h : ∀ b, k (I.fstTo b) ≫ I.fst b = k (I.sndTo b) ≫ I.snd b) : W ⟶ multiequalizer I := limit.lift _ (Multifork.ofι I _ k h) -@[reassoc] -- Porting note (#10618): simp can prove this, removed attribute +@[reassoc] theorem lift_ι (W : C) (k : ∀ a, W ⟶ I.left a) (h : ∀ b, k (I.fstTo b) ≫ I.fst b = k (I.sndTo b) ≫ I.snd b) (a) : Multiequalizer.lift I _ k h ≫ Multiequalizer.ι I a = k _ := @@ -792,7 +792,7 @@ abbrev desc (W : C) (k : ∀ b, I.right b ⟶ W) (h : ∀ a, I.fst a ≫ k (I.fstFrom a) = I.snd a ≫ k (I.sndFrom a)) : multicoequalizer I ⟶ W := colimit.desc _ (Multicofork.ofπ I _ k h) -@[reassoc] -- Porting note (#10618): simp can prove this, removed attribute +@[reassoc] theorem π_desc (W : C) (k : ∀ b, I.right b ⟶ W) (h : ∀ a, I.fst a ≫ k (I.fstFrom a) = I.snd a ≫ k (I.sndFrom a)) (b) : Multicoequalizer.π I b ≫ Multicoequalizer.desc I _ k h = k _ := diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Pullback/HasPullback.lean b/Mathlib/CategoryTheory/Limits/Shapes/Pullback/HasPullback.lean index f6455c517bc0b..7a97bde56f6a0 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Pullback/HasPullback.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Pullback/HasPullback.lean @@ -150,33 +150,27 @@ theorem PullbackCone.fst_limit_cone {X Y Z : C} (f : X ⟶ Z) (g : Y ⟶ Z) [Has theorem PullbackCone.snd_limit_cone {X Y Z : C} (f : X ⟶ Z) (g : Y ⟶ Z) [HasLimit (cospan f g)] : PullbackCone.snd (limit.cone (cospan f g)) = pullback.snd f g := rfl --- Porting note (#10618): simp can prove this; removed simp theorem PushoutCocone.inl_colimit_cocone {X Y Z : C} (f : Z ⟶ X) (g : Z ⟶ Y) [HasColimit (span f g)] : PushoutCocone.inl (colimit.cocone (span f g)) = pushout.inl _ _ := rfl --- Porting note (#10618): simp can prove this; removed simp theorem PushoutCocone.inr_colimit_cocone {X Y Z : C} (f : Z ⟶ X) (g : Z ⟶ Y) [HasColimit (span f g)] : PushoutCocone.inr (colimit.cocone (span f g)) = pushout.inr _ _ := rfl --- Porting note (#10618): simp can prove this and reassoced version; removed simp @[reassoc] theorem pullback.lift_fst {W X Y Z : C} {f : X ⟶ Z} {g : Y ⟶ Z} [HasPullback f g] (h : W ⟶ X) (k : W ⟶ Y) (w : h ≫ f = k ≫ g) : pullback.lift h k w ≫ pullback.fst f g = h := limit.lift_π _ _ --- Porting note (#10618): simp can prove this and reassoced version; removed simp @[reassoc] theorem pullback.lift_snd {W X Y Z : C} {f : X ⟶ Z} {g : Y ⟶ Z} [HasPullback f g] (h : W ⟶ X) (k : W ⟶ Y) (w : h ≫ f = k ≫ g) : pullback.lift h k w ≫ pullback.snd f g = k := limit.lift_π _ _ --- Porting note (#10618): simp can prove this and reassoced version; removed simp @[reassoc] theorem pushout.inl_desc {W X Y Z : C} {f : X ⟶ Y} {g : X ⟶ Z} [HasPushout f g] (h : Y ⟶ W) (k : Z ⟶ W) (w : f ≫ h = g ≫ k) : pushout.inl _ _ ≫ pushout.desc h k w = h := colimit.ι_desc _ _ --- Porting note (#10618): simp can prove this and reassoced version; removed simp @[reassoc] theorem pushout.inr_desc {W X Y Z : C} {f : X ⟶ Y} {g : X ⟶ Z} [HasPushout f g] (h : Y ⟶ W) (k : Z ⟶ W) (w : f ≫ h = g ≫ k) : pushout.inr _ _ ≫ pushout.desc h k w = k := diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Pullback/Iso.lean b/Mathlib/CategoryTheory/Limits/Shapes/Pullback/Iso.lean index 3007f243e923c..e55cb3b880356 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Pullback/Iso.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Pullback/Iso.lean @@ -42,7 +42,6 @@ theorem pullbackConeOfLeftIso_fst : (pullbackConeOfLeftIso f g).fst = g ≫ inv @[simp] theorem pullbackConeOfLeftIso_snd : (pullbackConeOfLeftIso f g).snd = 𝟙 _ := rfl --- Porting note (#10618): simp can prove this; removed simp theorem pullbackConeOfLeftIso_π_app_none : (pullbackConeOfLeftIso f g).π.app none = g := by simp @[simp] @@ -93,7 +92,6 @@ theorem pullbackConeOfRightIso_fst : (pullbackConeOfRightIso f g).fst = 𝟙 _ : @[simp] theorem pullbackConeOfRightIso_snd : (pullbackConeOfRightIso f g).snd = f ≫ inv g := rfl --- Porting note (#10618): simp can prove this; removed simps theorem pullbackConeOfRightIso_π_app_none : (pullbackConeOfRightIso f g).π.app none = f := by simp @[simp] @@ -145,7 +143,6 @@ theorem pushoutCoconeOfLeftIso_inl : (pushoutCoconeOfLeftIso f g).inl = inv f @[simp] theorem pushoutCoconeOfLeftIso_inr : (pushoutCoconeOfLeftIso f g).inr = 𝟙 _ := rfl --- Porting note (#10618): simp can prove this; removed simp theorem pushoutCoconeOfLeftIso_ι_app_none : (pushoutCoconeOfLeftIso f g).ι.app none = g := by simp @@ -197,7 +194,6 @@ theorem pushoutCoconeOfRightIso_inl : (pushoutCoconeOfRightIso f g).inl = 𝟙 _ @[simp] theorem pushoutCoconeOfRightIso_inr : (pushoutCoconeOfRightIso f g).inr = inv g ≫ f := rfl --- Porting note (#10618): simp can prove this; removed simp theorem pushoutCoconeOfRightIso_ι_app_none : (pushoutCoconeOfRightIso f g).ι.app none = f := by simp diff --git a/Mathlib/CategoryTheory/Limits/Shapes/WidePullbacks.lean b/Mathlib/CategoryTheory/Limits/Shapes/WidePullbacks.lean index 8b87766c91f50..bb674ca92f2a0 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/WidePullbacks.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/WidePullbacks.lean @@ -311,12 +311,10 @@ noncomputable abbrev lift {X : C} (f : X ⟶ B) (fs : ∀ j : J, X ⟶ objs j) variable (arrows) variable {X : C} (f : X ⟶ B) (fs : ∀ j : J, X ⟶ objs j) (w : ∀ j, fs j ≫ arrows j = f) --- Porting note (#10618): simp can prove this so removed simp attribute @[reassoc] theorem lift_π (j : J) : lift f fs w ≫ π arrows j = fs _ := by simp only [limit.lift_π, WidePullbackShape.mkCone_pt, WidePullbackShape.mkCone_π_app] --- Porting note (#10618): simp can prove this so removed simp attribute @[reassoc] theorem lift_base : lift f fs w ≫ base arrows = f := by simp only [limit.lift_π, WidePullbackShape.mkCone_pt, WidePullbackShape.mkCone_π_app] @@ -383,7 +381,6 @@ variable {X : C} (f : B ⟶ X) (fs : ∀ j : J, objs j ⟶ X) (w : ∀ j, arrows theorem ι_desc (j : J) : ι arrows j ≫ desc f fs w = fs _ := by simp only [colimit.ι_desc, WidePushoutShape.mkCocone_pt, WidePushoutShape.mkCocone_ι_app] --- Porting note (#10618): simp can prove this so removed simp attribute @[reassoc] theorem head_desc : head arrows ≫ desc f fs w = f := by simp only [colimit.ι_desc, WidePushoutShape.mkCocone_pt, WidePushoutShape.mkCocone_ι_app] diff --git a/Mathlib/CategoryTheory/PUnit.lean b/Mathlib/CategoryTheory/PUnit.lean index b561c063d4468..840d3c4759134 100644 --- a/Mathlib/CategoryTheory/PUnit.lean +++ b/Mathlib/CategoryTheory/PUnit.lean @@ -28,8 +28,6 @@ namespace Functor @[simps!] def star : C ⥤ Discrete PUnit.{w + 1} := (Functor.const _).obj ⟨⟨⟩⟩ --- Porting note (#10618): simp can simplify this -attribute [nolint simpNF] star_map_down_down variable {C} /-- Any two functors to `Discrete PUnit` are isomorphic. -/ diff --git a/Mathlib/CategoryTheory/Preadditive/InjectiveResolution.lean b/Mathlib/CategoryTheory/Preadditive/InjectiveResolution.lean index 9f756b5cdb943..8744f1317727f 100644 --- a/Mathlib/CategoryTheory/Preadditive/InjectiveResolution.lean +++ b/Mathlib/CategoryTheory/Preadditive/InjectiveResolution.lean @@ -92,13 +92,11 @@ lemma exact_succ (n : ℕ) : theorem ι_f_succ (n : ℕ) : I.ι.f (n + 1) = 0 := (isZero_single_obj_X _ _ _ _ (by simp)).eq_of_src _ _ --- Porting note (#10618): removed @[simp] simp can prove this @[reassoc] theorem ι_f_zero_comp_complex_d : I.ι.f 0 ≫ I.cocomplex.d 0 1 = 0 := by simp --- Porting note (#10618): removed @[simp] simp can prove this theorem complex_d_comp (n : ℕ) : I.cocomplex.d n (n + 1) ≫ I.cocomplex.d (n + 1) (n + 2) = 0 := by simp diff --git a/Mathlib/CategoryTheory/Preadditive/ProjectiveResolution.lean b/Mathlib/CategoryTheory/Preadditive/ProjectiveResolution.lean index 1447ce251a876..51a7ef3dae857 100644 --- a/Mathlib/CategoryTheory/Preadditive/ProjectiveResolution.lean +++ b/Mathlib/CategoryTheory/Preadditive/ProjectiveResolution.lean @@ -89,7 +89,6 @@ theorem complex_d_comp_π_f_zero : P.complex.d 1 0 ≫ P.π.f 0 = 0 := by rw [← P.π.comm 1 0, single_obj_d, comp_zero] --- Porting note (#10618): removed @[simp] simp can prove this theorem complex_d_succ_comp (n : ℕ) : P.complex.d n (n + 1) ≫ P.complex.d (n + 1) (n + 2) = 0 := by simp From 520965a9eec313a18fdc62faf79e4b875dab88d9 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Tue, 22 Oct 2024 15:31:19 +0000 Subject: [PATCH 282/505] chore: deprecate `WithTop.zero_lt_{coe,top}` for `WithTop.{coe,top}_pos` (#18048) The two are almost identical, except `zero_lt_coe` has an explicit argument where `coe_pos` has it implicit. Both the name and implicitness of `coe_pos` fit our style better. Also `zero_lt_coe` seems to be used only once in Mathlib. Found when cleaning up `simp` porting notes. --- Mathlib/Algebra/Order/Monoid/Unbundled/WithTop.lean | 12 ++++++------ .../Algebra/Polynomial/Degree/TrailingDegree.lean | 2 +- Mathlib/Analysis/Analytic/Constructions.lean | 2 +- Mathlib/Analysis/Normed/Algebra/Exponential.lean | 2 +- .../Analysis/SpecialFunctions/JapaneseBracket.lean | 2 +- Mathlib/Data/ENat/Basic.lean | 7 +++++-- .../Dynamics/TopologicalEntropy/CoverEntropy.lean | 2 +- Mathlib/MeasureTheory/Function/SimpleFunc.lean | 4 ++-- Mathlib/MeasureTheory/Measure/Count.lean | 4 ++-- Mathlib/Topology/MetricSpace/Infsep.lean | 2 +- 10 files changed, 21 insertions(+), 18 deletions(-) diff --git a/Mathlib/Algebra/Order/Monoid/Unbundled/WithTop.lean b/Mathlib/Algebra/Order/Monoid/Unbundled/WithTop.lean index 0b75be72f5aaf..b7cbbd3985728 100644 --- a/Mathlib/Algebra/Order/Monoid/Unbundled/WithTop.lean +++ b/Mathlib/Algebra/Order/Monoid/Unbundled/WithTop.lean @@ -368,13 +368,13 @@ instance existsAddOfLE [LE α] [Add α] [ExistsAddOfLE α] : ExistsAddOfLE (With -- CanonicallyLinearOrderedAddCommMonoid (WithTop α) := -- { WithTop.canonicallyOrderedAddCommMonoid, WithTop.linearOrder with } -@[simp] -theorem zero_lt_top [Zero α] [LT α] : (0 : WithTop α) < ⊤ := - coe_lt_top 0 +@[to_additive (attr := simp) top_pos] +theorem one_lt_top [One α] [LT α] : (1 : WithTop α) < ⊤ := coe_lt_top _ --- Porting note (#10618): simp can already prove this. --- @[simp] -@[norm_cast] +@[deprecated top_pos (since := "2024-10-22")] +alias zero_lt_top := top_pos + +@[norm_cast, deprecated coe_pos (since := "2024-10-22")] theorem zero_lt_coe [Zero α] [LT α] (a : α) : (0 : WithTop α) < a ↔ 0 < a := coe_lt_coe diff --git a/Mathlib/Algebra/Polynomial/Degree/TrailingDegree.lean b/Mathlib/Algebra/Polynomial/Degree/TrailingDegree.lean index 78ee6ce26e4a0..d68474f05ed6e 100644 --- a/Mathlib/Algebra/Polynomial/Degree/TrailingDegree.lean +++ b/Mathlib/Algebra/Polynomial/Degree/TrailingDegree.lean @@ -228,7 +228,7 @@ theorem coeff_eq_zero_of_lt_natTrailingDegree {p : R[X]} {n : ℕ} (h : n < p.na theorem coeff_natTrailingDegree_pred_eq_zero {p : R[X]} {hp : (0 : ℕ∞) < natTrailingDegree p} : p.coeff (p.natTrailingDegree - 1) = 0 := coeff_eq_zero_of_lt_natTrailingDegree <| - Nat.sub_lt ((WithTop.zero_lt_coe (natTrailingDegree p)).mp hp) Nat.one_pos + Nat.sub_lt (WithTop.coe_pos.mp hp) Nat.one_pos theorem le_trailingDegree_X_pow (n : ℕ) : (n : ℕ∞) ≤ trailingDegree (X ^ n : R[X]) := by simpa only [C_1, one_mul] using le_trailingDegree_C_mul_X_pow n (1 : R) diff --git a/Mathlib/Analysis/Analytic/Constructions.lean b/Mathlib/Analysis/Analytic/Constructions.lean index 61e423dab7c7d..c3c7fe59dc3a5 100644 --- a/Mathlib/Analysis/Analytic/Constructions.lean +++ b/Mathlib/Analysis/Analytic/Constructions.lean @@ -38,7 +38,7 @@ variable {A : Type*} [NormedRing A] [NormedAlgebra 𝕜 A] theorem hasFPowerSeriesOnBall_const {c : F} {e : E} : HasFPowerSeriesOnBall (fun _ => c) (constFormalMultilinearSeries 𝕜 E c) e ⊤ := by - refine ⟨by simp, WithTop.zero_lt_top, fun _ => hasSum_single 0 fun n hn => ?_⟩ + refine ⟨by simp, WithTop.top_pos, fun _ => hasSum_single 0 fun n hn => ?_⟩ simp [constFormalMultilinearSeries_apply hn] theorem hasFPowerSeriesAt_const {c : F} {e : E} : diff --git a/Mathlib/Analysis/Normed/Algebra/Exponential.lean b/Mathlib/Analysis/Normed/Algebra/Exponential.lean index 43b9083f92da7..a417ad5c90908 100644 --- a/Mathlib/Analysis/Normed/Algebra/Exponential.lean +++ b/Mathlib/Analysis/Normed/Algebra/Exponential.lean @@ -386,7 +386,7 @@ theorem expSeries_radius_eq_top : (expSeries 𝕂 𝔸).radius = ∞ := by theorem expSeries_radius_pos : 0 < (expSeries 𝕂 𝔸).radius := by rw [expSeries_radius_eq_top] - exact WithTop.zero_lt_top + exact WithTop.top_pos variable {𝕂 𝔸 𝔹} diff --git a/Mathlib/Analysis/SpecialFunctions/JapaneseBracket.lean b/Mathlib/Analysis/SpecialFunctions/JapaneseBracket.lean index 5b9cbab201ab2..52481419b2ca5 100644 --- a/Mathlib/Analysis/SpecialFunctions/JapaneseBracket.lean +++ b/Mathlib/Analysis/SpecialFunctions/JapaneseBracket.lean @@ -126,7 +126,7 @@ theorem finite_integral_one_add_norm {r : ℝ} (hnr : (finrank ℝ E : ℝ) < r) -- The integral over the constant zero function is finite: rw [setLIntegral_congr_fun measurableSet_Ioi (ae_of_all volume <| h_int''), lintegral_const 0, zero_mul] - exact WithTop.zero_lt_top + exact WithTop.top_pos theorem integrable_one_add_norm {r : ℝ} (hnr : (finrank ℝ E : ℝ) < r) : Integrable (fun x ↦ (1 + ‖x‖) ^ (-r)) μ := by diff --git a/Mathlib/Data/ENat/Basic.lean b/Mathlib/Data/ENat/Basic.lean index 3486a11a1cef2..02639995e5e57 100644 --- a/Mathlib/Data/ENat/Basic.lean +++ b/Mathlib/Data/ENat/Basic.lean @@ -175,8 +175,11 @@ theorem top_sub_ofNat (a : ℕ) [a.AtLeastTwo] : (⊤ : ℕ∞) - (no_index (OfN top_sub_coe a @[simp] -theorem zero_lt_top : (0 : ℕ∞) < ⊤ := - WithTop.zero_lt_top +theorem top_pos : (0 : ℕ∞) < ⊤ := + WithTop.top_pos + +@[deprecated ENat.top_pos (since := "2024-10-22")] +alias zero_lt_top := top_pos theorem sub_top (a : ℕ∞) : a - ⊤ = 0 := WithTop.sub_top diff --git a/Mathlib/Dynamics/TopologicalEntropy/CoverEntropy.lean b/Mathlib/Dynamics/TopologicalEntropy/CoverEntropy.lean index 2b8f40fded2bf..b5c4368d3eecd 100644 --- a/Mathlib/Dynamics/TopologicalEntropy/CoverEntropy.lean +++ b/Mathlib/Dynamics/TopologicalEntropy/CoverEntropy.lean @@ -274,7 +274,7 @@ lemma coverMincard_eq_zero_iff (T : X → X) (F : Set X) (U : Set (X × X)) (n : coverMincard T F U n = 0 ↔ F = ∅ := by refine Iff.intro (fun h ↦ subset_empty_iff.1 ?_) (fun F_empt ↦ by rw [F_empt, coverMincard_empty]) have := coverMincard_finite_iff T F U n - rw [h, eq_true ENat.zero_lt_top, true_iff] at this + rw [h, eq_true ENat.top_pos, true_iff] at this simp only [IsDynCoverOf, Finset.mem_coe, Nat.cast_eq_zero, Finset.card_eq_zero, exists_eq_right, Finset.not_mem_empty, iUnion_of_empty, iUnion_empty] at this exact this diff --git a/Mathlib/MeasureTheory/Function/SimpleFunc.lean b/Mathlib/MeasureTheory/Function/SimpleFunc.lean index 267c67d35ba52..6a1c98e568221 100644 --- a/Mathlib/MeasureTheory/Function/SimpleFunc.lean +++ b/Mathlib/MeasureTheory/Function/SimpleFunc.lean @@ -730,7 +730,7 @@ def eapprox : (α → ℝ≥0∞) → ℕ → α →ₛ ℝ≥0∞ := theorem eapprox_lt_top (f : α → ℝ≥0∞) (n : ℕ) (a : α) : eapprox f n a < ∞ := by simp only [eapprox, approx, finset_sup_apply, Finset.mem_range, ENNReal.bot_eq_zero, restrict] - rw [Finset.sup_lt_iff (α := ℝ≥0∞) WithTop.zero_lt_top] + rw [Finset.sup_lt_iff (α := ℝ≥0∞) WithTop.top_pos] intro b _ split_ifs · simp only [coe_zero, coe_piecewise, piecewise_eq_indicator, coe_const] @@ -739,7 +739,7 @@ theorem eapprox_lt_top (f : α → ℝ≥0∞) (n : ℕ) (a : α) : eapprox f n ennrealRatEmbed b := indicator_le_self _ _ a _ < ⊤ := ENNReal.coe_lt_top - · exact WithTop.zero_lt_top + · exact WithTop.top_pos @[mono] theorem monotone_eapprox (f : α → ℝ≥0∞) : Monotone (eapprox f) := diff --git a/Mathlib/MeasureTheory/Measure/Count.lean b/Mathlib/MeasureTheory/Measure/Count.lean index 75eef1a163c72..121403addfb48 100644 --- a/Mathlib/MeasureTheory/Measure/Count.lean +++ b/Mathlib/MeasureTheory/Measure/Count.lean @@ -102,13 +102,13 @@ theorem count_apply_lt_top [MeasurableSingletonClass α] : count s < ∞ ↔ s.F theorem empty_of_count_eq_zero' (s_mble : MeasurableSet s) (hsc : count s = 0) : s = ∅ := by have hs : s.Finite := by rw [← count_apply_lt_top' s_mble, hsc] - exact WithTop.zero_lt_top + exact WithTop.top_pos simpa [count_apply_finite' hs s_mble] using hsc theorem empty_of_count_eq_zero [MeasurableSingletonClass α] (hsc : count s = 0) : s = ∅ := by have hs : s.Finite := by rw [← count_apply_lt_top, hsc] - exact WithTop.zero_lt_top + exact WithTop.top_pos simpa [count_apply_finite _ hs] using hsc @[simp] diff --git a/Mathlib/Topology/MetricSpace/Infsep.lean b/Mathlib/Topology/MetricSpace/Infsep.lean index e451f9e7817b3..c47889bcd8b41 100644 --- a/Mathlib/Topology/MetricSpace/Infsep.lean +++ b/Mathlib/Topology/MetricSpace/Infsep.lean @@ -246,7 +246,7 @@ theorem einfsep_pos_of_finite [Finite s] : 0 < s.einfsep := by · rcases hs.einfsep_exists_of_finite with ⟨x, _hx, y, _hy, hxy, hxy'⟩ exact hxy'.symm ▸ edist_pos.2 hxy · rw [not_nontrivial_iff] at hs - exact hs.einfsep.symm ▸ WithTop.zero_lt_top + exact hs.einfsep.symm ▸ WithTop.top_pos theorem relatively_discrete_of_finite [Finite s] : ∃ C > 0, ∀ x ∈ s, ∀ y ∈ s, x ≠ y → C ≤ edist x y := by From e7dbdb48a63351e003b88001e2371ab4f0a24243 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Tue, 22 Oct 2024 15:31:20 +0000 Subject: [PATCH 283/505] =?UTF-8?q?chore:=20use=20new=20`=E2=84=B5=5F`=20a?= =?UTF-8?q?nd=20`=E2=84=B6=5F`=20notation=20(#18066)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- Counterexamples/Phillips.lean | 26 +++++++++++----------- Mathlib/Order/Filter/CardinalInter.lean | 20 ++++++++--------- Mathlib/Order/Filter/CountableInter.lean | 4 ++-- Mathlib/SetTheory/Cardinal/Aleph.lean | 4 ++-- Mathlib/SetTheory/Cardinal/Arithmetic.lean | 10 ++++----- Mathlib/SetTheory/Cardinal/Cofinality.lean | 8 +++---- 6 files changed, 36 insertions(+), 36 deletions(-) diff --git a/Counterexamples/Phillips.lean b/Counterexamples/Phillips.lean index cf03d0dc7c250..b1bd199ed054e 100644 --- a/Counterexamples/Phillips.lean +++ b/Counterexamples/Phillips.lean @@ -461,7 +461,7 @@ We need the continuum hypothesis to construct it. -/ -theorem sierpinski_pathological_family (Hcont : #ℝ = aleph 1) : +theorem sierpinski_pathological_family (Hcont : #ℝ = ℵ₁) : ∃ f : ℝ → Set ℝ, (∀ x, (univ \ f x).Countable) ∧ ∀ y, {x : ℝ | y ∈ f x}.Countable := by rcases Cardinal.ord_eq ℝ with ⟨r, hr, H⟩ refine ⟨fun x => {y | r x y}, fun x => ?_, fun y => ?_⟩ @@ -484,13 +484,13 @@ theorem sierpinski_pathological_family (Hcont : #ℝ = aleph 1) : /-- A family of sets in `ℝ` which only miss countably many points, but such that any point is contained in only countably many of them. -/ -def spf (Hcont : #ℝ = aleph 1) (x : ℝ) : Set ℝ := +def spf (Hcont : #ℝ = ℵ₁) (x : ℝ) : Set ℝ := (sierpinski_pathological_family Hcont).choose x -theorem countable_compl_spf (Hcont : #ℝ = aleph 1) (x : ℝ) : (univ \ spf Hcont x).Countable := +theorem countable_compl_spf (Hcont : #ℝ = ℵ₁) (x : ℝ) : (univ \ spf Hcont x).Countable := (sierpinski_pathological_family Hcont).choose_spec.1 x -theorem countable_spf_mem (Hcont : #ℝ = aleph 1) (y : ℝ) : {x | y ∈ spf Hcont x}.Countable := +theorem countable_spf_mem (Hcont : #ℝ = ℵ₁) (y : ℝ) : {x | y ∈ spf Hcont x}.Countable := (sierpinski_pathological_family Hcont).choose_spec.2 y /-! @@ -510,10 +510,10 @@ which is large (it has countable complement), as in the Sierpinski pathological /-- A family of bounded functions `f_x` from `ℝ` (seen with the discrete topology) to `ℝ` (in fact taking values in `{0, 1}`), indexed by a real parameter `x`, corresponding to the characteristic functions of the different fibers of the Sierpinski pathological family -/ -def f (Hcont : #ℝ = aleph 1) (x : ℝ) : DiscreteCopy ℝ →ᵇ ℝ := +def f (Hcont : #ℝ = ℵ₁) (x : ℝ) : DiscreteCopy ℝ →ᵇ ℝ := ofNormedAddCommGroupDiscrete (indicator (spf Hcont x) 1) 1 (norm_indicator_le_one _) -theorem apply_f_eq_continuousPart (Hcont : #ℝ = aleph 1) (φ : (DiscreteCopy ℝ →ᵇ ℝ) →L[ℝ] ℝ) +theorem apply_f_eq_continuousPart (Hcont : #ℝ = ℵ₁) (φ : (DiscreteCopy ℝ →ᵇ ℝ) →L[ℝ] ℝ) (x : ℝ) (hx : φ.toBoundedAdditiveMeasure.discreteSupport ∩ spf Hcont x = ∅) : φ (f Hcont x) = φ.toBoundedAdditiveMeasure.continuousPart univ := by set ψ := φ.toBoundedAdditiveMeasure @@ -523,7 +523,7 @@ theorem apply_f_eq_continuousPart (Hcont : #ℝ = aleph 1) (φ : (DiscreteCopy ψ.continuousPart.additive _ _ disjoint_sdiff_self_right, ψ.continuousPart_apply_eq_zero_of_countable _ (countable_compl_spf Hcont x), add_zero] -theorem countable_ne (Hcont : #ℝ = aleph 1) (φ : (DiscreteCopy ℝ →ᵇ ℝ) →L[ℝ] ℝ) : +theorem countable_ne (Hcont : #ℝ = ℵ₁) (φ : (DiscreteCopy ℝ →ᵇ ℝ) →L[ℝ] ℝ) : {x | φ.toBoundedAdditiveMeasure.continuousPart univ ≠ φ (f Hcont x)}.Countable := by have A : {x | φ.toBoundedAdditiveMeasure.continuousPart univ ≠ φ (f Hcont x)} ⊆ @@ -543,7 +543,7 @@ theorem countable_ne (Hcont : #ℝ = aleph 1) (φ : (DiscreteCopy ℝ →ᵇ ℝ apply Countable.mono (Subset.trans A B) exact Countable.biUnion (countable_discreteSupport _) fun a _ => countable_spf_mem Hcont a -theorem comp_ae_eq_const (Hcont : #ℝ = aleph 1) (φ : (DiscreteCopy ℝ →ᵇ ℝ) →L[ℝ] ℝ) : +theorem comp_ae_eq_const (Hcont : #ℝ = ℵ₁) (φ : (DiscreteCopy ℝ →ᵇ ℝ) →L[ℝ] ℝ) : ∀ᵐ x ∂volume.restrict (Icc (0 : ℝ) 1), φ.toBoundedAdditiveMeasure.continuousPart univ = φ (f Hcont x) := by apply ae_restrict_of_ae @@ -551,7 +551,7 @@ theorem comp_ae_eq_const (Hcont : #ℝ = aleph 1) (φ : (DiscreteCopy ℝ →ᵇ intro x simp only [imp_self, mem_setOf_eq, mem_compl_iff] -theorem integrable_comp (Hcont : #ℝ = aleph 1) (φ : (DiscreteCopy ℝ →ᵇ ℝ) →L[ℝ] ℝ) : +theorem integrable_comp (Hcont : #ℝ = ℵ₁) (φ : (DiscreteCopy ℝ →ᵇ ℝ) →L[ℝ] ℝ) : IntegrableOn (fun x => φ (f Hcont x)) (Icc 0 1) := by have : IntegrableOn (fun _ => φ.toBoundedAdditiveMeasure.continuousPart univ) (Icc (0 : ℝ) 1) @@ -559,7 +559,7 @@ theorem integrable_comp (Hcont : #ℝ = aleph 1) (φ : (DiscreteCopy ℝ →ᵇ simp [integrableOn_const] apply Integrable.congr this (comp_ae_eq_const Hcont φ) -theorem integral_comp (Hcont : #ℝ = aleph 1) (φ : (DiscreteCopy ℝ →ᵇ ℝ) →L[ℝ] ℝ) : +theorem integral_comp (Hcont : #ℝ = ℵ₁) (φ : (DiscreteCopy ℝ →ᵇ ℝ) →L[ℝ] ℝ) : ∫ x in Icc 0 1, φ (f Hcont x) = φ.toBoundedAdditiveMeasure.continuousPart univ := by rw [← integral_congr_ae (comp_ae_eq_const Hcont φ)] simp @@ -574,18 +574,18 @@ no Pettis integral. example : CompleteSpace (DiscreteCopy ℝ →ᵇ ℝ) := by infer_instance /-- The function `f Hcont : ℝ → (DiscreteCopy ℝ →ᵇ ℝ)` is scalarly measurable. -/ -theorem measurable_comp (Hcont : #ℝ = aleph 1) (φ : (DiscreteCopy ℝ →ᵇ ℝ) →L[ℝ] ℝ) : +theorem measurable_comp (Hcont : #ℝ = ℵ₁) (φ : (DiscreteCopy ℝ →ᵇ ℝ) →L[ℝ] ℝ) : Measurable fun x => φ (f Hcont x) := by have : Measurable fun _ : ℝ => φ.toBoundedAdditiveMeasure.continuousPart univ := measurable_const refine this.measurable_of_countable_ne ?_ exact countable_ne Hcont φ /-- The function `f Hcont : ℝ → (DiscreteCopy ℝ →ᵇ ℝ)` is uniformly bounded by `1` in norm. -/ -theorem norm_bound (Hcont : #ℝ = aleph 1) (x : ℝ) : ‖f Hcont x‖ ≤ 1 := +theorem norm_bound (Hcont : #ℝ = ℵ₁) (x : ℝ) : ‖f Hcont x‖ ≤ 1 := norm_ofNormedAddCommGroup_le _ zero_le_one (norm_indicator_le_one _) /-- The function `f Hcont : ℝ → (DiscreteCopy ℝ →ᵇ ℝ)` has no Pettis integral. -/ -theorem no_pettis_integral (Hcont : #ℝ = aleph 1) : +theorem no_pettis_integral (Hcont : #ℝ = ℵ₁) : ¬∃ g : DiscreteCopy ℝ →ᵇ ℝ, ∀ φ : (DiscreteCopy ℝ →ᵇ ℝ) →L[ℝ] ℝ, ∫ x in Icc 0 1, φ (f Hcont x) = φ g := by rintro ⟨g, h⟩ diff --git a/Mathlib/Order/Filter/CardinalInter.lean b/Mathlib/Order/Filter/CardinalInter.lean index bfce10f3629f3..22967d8a816b3 100644 --- a/Mathlib/Order/Filter/CardinalInter.lean +++ b/Mathlib/Order/Filter/CardinalInter.lean @@ -17,11 +17,11 @@ their intersection belongs to `l` as well. # Main results * `Filter.cardinalInterFilter_aleph0` establishes that every filter `l` is a - `CardinalInterFilter l aleph0` + `CardinalInterFilter l ℵ₀` * `CardinalInterFilter.toCountableInterFilter` establishes that every `CardinalInterFilter l c` with - `c > aleph0` is a `CountableInterFilter`. + `c > ℵ₀` is a `CountableInterFilter`. * `CountableInterFilter.toCardinalInterFilter` establishes that every `CountableInterFilter l` is a - `CardinalInterFilter l aleph1`. + `CardinalInterFilter l ℵ₁`. * `CardinalInterFilter.of_CardinalInterFilter_of_lt` establishes that we have `CardinalInterFilter l c` → `CardinalInterFilter l a` for all `a < c`. @@ -48,26 +48,26 @@ theorem cardinal_sInter_mem {S : Set (Set α)} [CardinalInterFilter l c] (hSc : ⋂₀ S ∈ l ↔ ∀ s ∈ S, s ∈ l := ⟨fun hS _s hs => mem_of_superset hS (sInter_subset_of_mem hs), CardinalInterFilter.cardinal_sInter_mem _ hSc⟩ -/-- Every filter is a CardinalInterFilter with c = aleph0 -/ -theorem _root_.Filter.cardinalInterFilter_aleph0 (l : Filter α) : CardinalInterFilter l aleph0 where +/-- Every filter is a CardinalInterFilter with c = ℵ₀ -/ +theorem _root_.Filter.cardinalInterFilter_aleph0 (l : Filter α) : CardinalInterFilter l ℵ₀ where cardinal_sInter_mem := by simp_all only [aleph_zero, lt_aleph0_iff_subtype_finite, setOf_mem_eq, sInter_mem, implies_true, forall_const] -/-- Every CardinalInterFilter with c > aleph0 is a CountableInterFilter -/ +/-- Every CardinalInterFilter with c > ℵ₀ is a CountableInterFilter -/ theorem CardinalInterFilter.toCountableInterFilter (l : Filter α) [CardinalInterFilter l c] - (hc : aleph0 < c) : CountableInterFilter l where + (hc : ℵ₀ < c) : CountableInterFilter l where countable_sInter_mem S hS a := CardinalInterFilter.cardinal_sInter_mem S (lt_of_le_of_lt (Set.Countable.le_aleph0 hS) hc) a -/-- Every CountableInterFilter is a CardinalInterFilter with c = aleph 1-/ +/-- Every CountableInterFilter is a CardinalInterFilter with c = ℵ₁ -/ instance CountableInterFilter.toCardinalInterFilter (l : Filter α) [CountableInterFilter l] : - CardinalInterFilter l (aleph 1) where + CardinalInterFilter l ℵ₁ where cardinal_sInter_mem S hS a := CountableInterFilter.countable_sInter_mem S ((countable_iff_lt_aleph_one S).mpr hS) a theorem cardinalInterFilter_aleph_one_iff : - CardinalInterFilter l (aleph 1) ↔ CountableInterFilter l := + CardinalInterFilter l ℵ₁ ↔ CountableInterFilter l := ⟨fun _ ↦ ⟨fun S h a ↦ CardinalInterFilter.cardinal_sInter_mem S ((countable_iff_lt_aleph_one S).1 h) a⟩, fun _ ↦ CountableInterFilter.toCardinalInterFilter l⟩ diff --git a/Mathlib/Order/Filter/CountableInter.lean b/Mathlib/Order/Filter/CountableInter.lean index b9d5d1ef60f8d..829b92767118e 100644 --- a/Mathlib/Order/Filter/CountableInter.lean +++ b/Mathlib/Order/Filter/CountableInter.lean @@ -21,8 +21,8 @@ and provide instances for some basic constructions (`⊥`, `⊤`, `Filter.princi that deduces two axioms of a `Filter` from the countable intersection property. Note that there also exists a typeclass `CardinalInterFilter`, and thus an alternative spelling of -`CountableInterFilter` as `CardinalInterFilter l (aleph 1)`. The former (defined here) is the -preferred spelling; it has the advantage of not requiring the user to import the theory ordinals. +`CountableInterFilter` as `CardinalInterFilter l ℵ₁`. The former (defined here) is the +preferred spelling; it has the advantage of not requiring the user to import the theory of ordinals. ## Tags filter, countable diff --git a/Mathlib/SetTheory/Cardinal/Aleph.lean b/Mathlib/SetTheory/Cardinal/Aleph.lean index d41e56c312df0..c1473e55bb543 100644 --- a/Mathlib/SetTheory/Cardinal/Aleph.lean +++ b/Mathlib/SetTheory/Cardinal/Aleph.lean @@ -321,13 +321,13 @@ theorem aleph_limit {o : Ordinal} (ho : o.IsLimit) : ℵ_ o = ⨆ a : Iio o, ℵ theorem aleph0_le_aleph' {o : Ordinal} : ℵ₀ ≤ aleph' o ↔ ω ≤ o := by rw [← aleph'_omega0, aleph'_le] -theorem aleph0_le_aleph (o : Ordinal) : ℵ₀ ≤ aleph o := by +theorem aleph0_le_aleph (o : Ordinal) : ℵ₀ ≤ ℵ_ o := by rw [aleph_eq_aleph', aleph0_le_aleph'] apply Ordinal.le_add_right theorem aleph'_pos {o : Ordinal} (ho : 0 < o) : 0 < aleph' o := by rwa [← aleph'_zero, aleph'_lt] -theorem aleph_pos (o : Ordinal) : 0 < aleph o := +theorem aleph_pos (o : Ordinal) : 0 < ℵ_ o := aleph0_pos.trans_le (aleph0_le_aleph o) @[simp] diff --git a/Mathlib/SetTheory/Cardinal/Arithmetic.lean b/Mathlib/SetTheory/Cardinal/Arithmetic.lean index 4dfd399fb40da..b942a95f09daa 100644 --- a/Mathlib/SetTheory/Cardinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Cardinal/Arithmetic.lean @@ -100,7 +100,7 @@ theorem mul_mk_eq_max {α β : Type u} [Infinite α] [Infinite β] : #α * #β = mul_eq_max (aleph0_le_mk α) (aleph0_le_mk β) @[simp] -theorem aleph_mul_aleph (o₁ o₂ : Ordinal) : aleph o₁ * aleph o₂ = aleph (max o₁ o₂) := by +theorem aleph_mul_aleph (o₁ o₂ : Ordinal) : ℵ_ o₁ * ℵ_ o₂ = ℵ_ (max o₁ o₂) := by rw [Cardinal.mul_eq_max (aleph0_le_aleph o₁) (aleph0_le_aleph o₂), aleph_max] @[simp] @@ -120,11 +120,11 @@ theorem mk_mul_aleph0_eq {α : Type*} [Infinite α] : #α * ℵ₀ = #α := mul_aleph0_eq (aleph0_le_mk α) @[simp] -theorem aleph0_mul_aleph (o : Ordinal) : ℵ₀ * aleph o = aleph o := +theorem aleph0_mul_aleph (o : Ordinal) : ℵ₀ * ℵ_ o = ℵ_ o := aleph0_mul_eq (aleph0_le_aleph o) @[simp] -theorem aleph_mul_aleph0 (o : Ordinal) : aleph o * ℵ₀ = aleph o := +theorem aleph_mul_aleph0 (o : Ordinal) : ℵ_ o * ℵ₀ = ℵ_ o := mul_aleph0_eq (aleph0_le_aleph o) theorem mul_lt_of_lt {a b c : Cardinal} (hc : ℵ₀ ≤ c) (h1 : a < c) (h2 : b < c) : a * b < c := @@ -421,7 +421,7 @@ end ciSup section aleph @[simp] -theorem aleph_add_aleph (o₁ o₂ : Ordinal) : aleph o₁ + aleph o₂ = aleph (max o₁ o₂) := by +theorem aleph_add_aleph (o₁ o₂ : Ordinal) : ℵ_ o₁ + ℵ_ o₂ = ℵ_ (max o₁ o₂) := by rw [Cardinal.add_eq_max (aleph0_le_aleph o₁), aleph_max] theorem principal_add_ord {c : Cardinal} (hc : ℵ₀ ≤ c) : Ordinal.Principal (· + ·) c.ord := @@ -429,7 +429,7 @@ theorem principal_add_ord {c : Cardinal} (hc : ℵ₀ ≤ c) : Ordinal.Principal rw [lt_ord, Ordinal.card_add] at * exact add_lt_of_lt hc ha hb -theorem principal_add_aleph (o : Ordinal) : Ordinal.Principal (· + ·) (aleph o).ord := +theorem principal_add_aleph (o : Ordinal) : Ordinal.Principal (· + ·) (ℵ_ o).ord := principal_add_ord <| aleph0_le_aleph o theorem add_right_inj_of_lt_aleph0 {α β γ : Cardinal} (γ₀ : γ < aleph0) : α + γ = β + γ ↔ α = β := diff --git a/Mathlib/SetTheory/Cardinal/Cofinality.lean b/Mathlib/SetTheory/Cardinal/Cofinality.lean index cbf48de84b877..e7fb33c1cf155 100644 --- a/Mathlib/SetTheory/Cardinal/Cofinality.lean +++ b/Mathlib/SetTheory/Cardinal/Cofinality.lean @@ -676,7 +676,7 @@ theorem aleph'_cof {o : Ordinal} (ho : o.IsLimit) : (aleph' o).ord.cof = o.cof : aleph'_isNormal.cof_eq ho @[simp] -theorem aleph_cof {o : Ordinal} (ho : o.IsLimit) : (aleph o).ord.cof = o.cof := +theorem aleph_cof {o : Ordinal} (ho : o.IsLimit) : (ℵ_ o).ord.cof = o.cof := aleph_isNormal.cof_eq ho @[simp] @@ -825,7 +825,7 @@ set_option linter.deprecated false in theorem IsStrongLimit.isLimit {c} (H : IsStrongLimit c) : IsLimit c := ⟨H.ne_zero, H.isSuccPrelimit⟩ -theorem isStrongLimit_beth {o : Ordinal} (H : IsSuccPrelimit o) : IsStrongLimit (beth o) := by +theorem isStrongLimit_beth {o : Ordinal} (H : IsSuccPrelimit o) : IsStrongLimit (ℶ_ o) := by rcases eq_or_ne o 0 with (rfl | h) · rw [beth_zero] exact isStrongLimit_aleph0 @@ -932,7 +932,7 @@ theorem isRegular_succ {c : Cardinal.{u}} (h : ℵ₀ ≤ c) : IsRegular (succ c · rw [← lt_succ_iff, ← lt_ord, ← αe, re] apply typein_lt_type)⟩ -theorem isRegular_aleph_one : IsRegular (aleph 1) := by +theorem isRegular_aleph_one : IsRegular ℵ₁ := by rw [← succ_aleph0] exact isRegular_succ le_rfl @@ -940,7 +940,7 @@ theorem isRegular_aleph'_succ {o : Ordinal} (h : ω ≤ o) : IsRegular (aleph' ( rw [aleph'_succ] exact isRegular_succ (aleph0_le_aleph'.2 h) -theorem isRegular_aleph_succ (o : Ordinal) : IsRegular (aleph (succ o)) := by +theorem isRegular_aleph_succ (o : Ordinal) : IsRegular (ℵ_ (succ o)) := by rw [aleph_succ] exact isRegular_succ (aleph0_le_aleph o) From c6021b7b74fdf03681a73e4bbcb7e6eee9f15ffa Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 22 Oct 2024 16:09:30 +0000 Subject: [PATCH 284/505] refactor(GroupAction/Blocks): additivise, fix names, golf API (#17578) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Most of the API can be additivised smoothly. The lemma names didn't match the naming convention. Some proofs were harder than necessary because of the suboptimal definition of `IsBlock` as `∀ _, _ ∨ _` rather than `∀ _, _ → _`. From LeanCamCombi --- Mathlib/Data/Set/Pairwise/Basic.lean | 8 +- Mathlib/GroupTheory/GroupAction/Basic.lean | 2 + Mathlib/GroupTheory/GroupAction/Blocks.lean | 476 ++++++++++---------- 3 files changed, 233 insertions(+), 253 deletions(-) diff --git a/Mathlib/Data/Set/Pairwise/Basic.lean b/Mathlib/Data/Set/Pairwise/Basic.lean index 7365c110ac476..ec5b714ca6538 100644 --- a/Mathlib/Data/Set/Pairwise/Basic.lean +++ b/Mathlib/Data/Set/Pairwise/Basic.lean @@ -296,12 +296,8 @@ lemma PairwiseDisjoint.eq_or_disjoint exact h.elim hi hj lemma pairwiseDisjoint_range_iff {α β : Type*} {f : α → (Set β)} : - (Set.range f).PairwiseDisjoint id ↔ ∀ x y, f x = f y ∨ Disjoint (f x) (f y) := by - constructor - · intro h x y - apply h.eq_or_disjoint (Set.mem_range_self x) (Set.mem_range_self y) - · rintro h _ ⟨x, rfl⟩ _ ⟨y, rfl⟩ hxy - exact (h x y).resolve_left hxy + (range f).PairwiseDisjoint id ↔ ∀ x y, f x ≠ f y → Disjoint (f x) (f y) := by + aesop (add simp [PairwiseDisjoint, Set.Pairwise]) /-- If the range of `f` is pairwise disjoint, then the image of any set `s` under `f` is as well. -/ lemma _root_.Pairwise.pairwiseDisjoint (h : Pairwise (Disjoint on f)) (s : Set ι) : diff --git a/Mathlib/GroupTheory/GroupAction/Basic.lean b/Mathlib/GroupTheory/GroupAction/Basic.lean index 547d1fc8f93c9..e890d7dc7383d 100644 --- a/Mathlib/GroupTheory/GroupAction/Basic.lean +++ b/Mathlib/GroupTheory/GroupAction/Basic.lean @@ -296,6 +296,8 @@ variable {G α β : Type*} [Group G] [MulAction G α] [MulAction G β] section Orbit +-- TODO: This proof is redoing a special case of `MulAction.IsInvariantBlock.isBlock`. Can we move +-- this lemma earlier to golf? @[to_additive (attr := simp)] theorem smul_orbit (g : G) (a : α) : g • orbit G a = orbit G a := (smul_orbit_subset g a).antisymm <| diff --git a/Mathlib/GroupTheory/GroupAction/Blocks.lean b/Mathlib/GroupTheory/GroupAction/Blocks.lean index cc63c90b25d16..054bfb7524698 100644 --- a/Mathlib/GroupTheory/GroupAction/Blocks.lean +++ b/Mathlib/GroupTheory/GroupAction/Blocks.lean @@ -3,11 +3,7 @@ Copyright (c) 2024 Antoine Chambert-Loir. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Antoine Chambert-Loir -/ - -import Mathlib.Algebra.Group.Subgroup.Actions -import Mathlib.Data.Set.Card import Mathlib.Data.Setoid.Partition -import Mathlib.GroupTheory.GroupAction.Basic import Mathlib.GroupTheory.GroupAction.Pointwise import Mathlib.GroupTheory.GroupAction.SubMulAction import Mathlib.GroupTheory.Index @@ -39,10 +35,10 @@ that contain a given point is a block ## References -We follow [wielandt1964]. - +We follow [Wielandt-1964]. -/ +open Set open scoped Pointwise namespace MulAction @@ -51,12 +47,14 @@ section orbits variable {G : Type*} [Group G] {X : Type*} [MulAction G X] +@[to_additive] theorem orbit.eq_or_disjoint (a b : X) : orbit G a = orbit G b ∨ Disjoint (orbit G a) (orbit G b) := by apply (em (Disjoint (orbit G a) (orbit G b))).symm.imp _ id simp (config := { contextual := true }) only [Set.not_disjoint_iff, ← orbit_eq_iff, forall_exists_index, and_imp, eq_comm, implies_true] +@[to_additive] theorem orbit.pairwiseDisjoint : (Set.range fun x : X => orbit G x).PairwiseDisjoint id := by rintro s ⟨x, rfl⟩ t ⟨y, rfl⟩ h @@ -64,6 +62,7 @@ theorem orbit.pairwiseDisjoint : exact (orbit.eq_or_disjoint x y).resolve_right h /-- Orbits of an element form a partition -/ +@[to_additive] theorem IsPartition.of_orbits : Setoid.IsPartition (Set.range fun a : X => orbit G a) := by apply orbit.pairwiseDisjoint.isPartition_of_exists_of_ne_empty @@ -76,246 +75,201 @@ end orbits section SMul -variable (G : Type*) {X : Type*} [SMul G X] +variable (G : Type*) {X : Type*} [SMul G X] {B : Set X} {a : X} -- Change terminology : is_fully_invariant ? -/-- For `SMul G X`, a fixed block is a `Set X` which is fully invariant: - `g • B = B` for all `g : G` -/ +/-- A set `B` is a `G`-fixed block if `g • B = B` for all `g : G`. -/ +@[to_additive "A set `B` is a `G`-fixed block if `g +ᵥ B = B` for all `g : G`."] def IsFixedBlock (B : Set X) := ∀ g : G, g • B = B -/-- For `SMul G X`, an invariant block is a `Set X` which is stable: - `g • B ⊆ B` for all `g : G` -/ +/-- A set `B` is a `G`-invariant block if `g • B ⊆ B` for all `g : G`. + +Note: It is not necessarily a block when the action is not by a group. -/ +@[to_additive +"A set `B` is a `G`-invariant block if `g +ᵥ B ⊆ B` for all `g : G`. + +Note: It is not necessarily a block when the action is not by a group. "] def IsInvariantBlock (B : Set X) := ∀ g : G, g • B ⊆ B -/-- A trivial block is a `Set X` which is either a subsingleton or ⊤ - (it is not necessarily a block…) -/ -def IsTrivialBlock (B : Set X) := B.Subsingleton ∨ B = ⊤ +/-- A trivial block is a `Set X` which is either a subsingleton or `univ`. -/-- `For SMul G X`, a block is a `Set X` whose translates are pairwise disjoint -/ -def IsBlock (B : Set X) := (Set.range fun g : G => g • B).PairwiseDisjoint id +Note: It is not necessarily a block when the action is not by a group. -/ +@[to_additive +"A trivial block is a `Set X` which is either a subsingleton or `univ`. + +Note: It is not necessarily a block when the action is not by a group."] +def IsTrivialBlock (B : Set X) := B.Subsingleton ∨ B = univ + +/-- A set `B` is a `G`-block iff the sets of the form `g • B` are pairwise equal or disjoint. -/ +@[to_additive +"A set `B` is a `G`-block iff the sets of the form `g +ᵥ B` are pairwise equal or disjoint. "] +def IsBlock (B : Set X) := ∀ ⦃g₁ g₂ : G⦄, g₁ • B ≠ g₂ • B → Disjoint (g₁ • B) (g₂ • B) variable {G} -/-- A set B is a block iff for all g, g', -the sets g • B and g' • B are either equal or disjoint -/ -theorem IsBlock.def {B : Set X} : - IsBlock G B ↔ ∀ g g' : G, g • B = g' • B ∨ Disjoint (g • B) (g' • B) := by - apply Set.pairwiseDisjoint_range_iff +@[to_additive] +lemma isBlock_iff_smul_eq_smul_of_nonempty : + IsBlock G B ↔ ∀ ⦃g₁ g₂ : G⦄, (g₁ • B ∩ g₂ • B).Nonempty → g₁ • B = g₂ • B := by + simp_rw [IsBlock, ← not_disjoint_iff_nonempty_inter, not_imp_comm] + +@[to_additive] +lemma isBlock_iff_pairwiseDisjoint_range_smul : + IsBlock G B ↔ (range fun g : G ↦ g • B).PairwiseDisjoint id := pairwiseDisjoint_range_iff.symm -/-- Alternate definition of a block -/ -theorem IsBlock.mk_notempty {B : Set X} : - IsBlock G B ↔ ∀ g g' : G, g • B ∩ g' • B ≠ ∅ → g • B = g' • B := by - simp_rw [IsBlock.def, or_iff_not_imp_right, Set.disjoint_iff_inter_eq_empty] +@[to_additive] +lemma isBlock_iff_smul_eq_smul_or_disjoint : + IsBlock G B ↔ ∀ g₁ g₂ : G, g₁ • B = g₂ • B ∨ Disjoint (g₁ • B) (g₂ • B) := + forall₂_congr fun _ _ ↦ or_iff_not_imp_left.symm -/-- A fixed block is a block -/ -theorem IsFixedBlock.isBlock {B : Set X} (hfB : IsFixedBlock G B) : - IsBlock G B := by - simp [IsBlock.def, hfB _] +alias ⟨IsBlock.smul_eq_smul_of_nonempty, _⟩ := isBlock_iff_smul_eq_smul_of_nonempty +alias ⟨IsBlock.pairwiseDisjoint_range_smul, _⟩ := isBlock_iff_pairwiseDisjoint_range_smul +alias ⟨IsBlock.smul_eq_smul_or_disjoint, _⟩ := isBlock_iff_smul_eq_smul_or_disjoint -variable (X) +/-- A fixed block is a block. -/ +@[to_additive "A fixed block is a block."] +lemma IsFixedBlock.isBlock (hfB : IsFixedBlock G B) : IsBlock G B := by simp [IsBlock, hfB _] -/-- The empty set is a block -/ -theorem isBlock_empty : IsBlock G (⊥ : Set X) := by - simp [IsBlock.def, Set.bot_eq_empty, Set.smul_set_empty] +/-- The empty set is a block. -/ +@[to_additive (attr := simp) "The empty set is a block."] +lemma IsBlock.empty : IsBlock G (∅ : Set X) := by simp [IsBlock] -variable {X} +/-- A singleton is a block. -/ +@[to_additive "A singleton is a block."] +lemma IsBlock.singleton : IsBlock G ({a} : Set X) := by simp [IsBlock] -theorem isBlock_singleton (a : X) : IsBlock G ({a} : Set X) := by - simp [IsBlock.def, Classical.or_iff_not_imp_left] +/-- Subsingletons are (trivial) blocks. -/ +@[to_additive "Subsingletons are (trivial) blocks."] +lemma IsBlock.of_subsingleton (hB : B.Subsingleton) : IsBlock G B := + hB.induction_on .empty fun _ ↦ .singleton -/-- Subsingletons are (trivial) blocks -/ -theorem isBlock_subsingleton {B : Set X} (hB : B.Subsingleton) : - IsBlock G B := - hB.induction_on (isBlock_empty _) isBlock_singleton +/-- A fixed block is an invariant block. -/ +@[to_additive "A fixed block is an invariant block."] +lemma IsFixedBlock.isInvariantBlock (hB : IsFixedBlock G B) : IsInvariantBlock G B := + fun _ ↦ (hB _).le end SMul section Group -variable {G : Type*} [Group G] {X : Type*} [MulAction G X] +variable {G : Type*} [Group G] {X : Type*} [MulAction G X] {B : Set X} + +@[to_additive] +lemma isBlock_iff_disjoint_smul_of_ne : + IsBlock G B ↔ ∀ ⦃g : G⦄, g • B ≠ B → Disjoint (g • B) B := by + refine ⟨fun hB g ↦ by simpa using hB (g₂ := 1), fun hB g₁ g₂ h ↦ ?_⟩ + simp only [disjoint_smul_set_right, ne_eq, ← inv_smul_eq_iff, smul_smul] at h ⊢ + exact hB h + +@[to_additive] +lemma isBlock_iff_smul_eq_of_nonempty : + IsBlock G B ↔ ∀ ⦃g : G⦄, (g • B ∩ B).Nonempty → g • B = B := by + simp_rw [isBlock_iff_disjoint_smul_of_ne, ← not_disjoint_iff_nonempty_inter, not_imp_comm] + +@[to_additive] +lemma isBlock_iff_smul_eq_or_disjoint : + IsBlock G B ↔ ∀ g : G, g • B = B ∨ Disjoint (g • B) B := + isBlock_iff_disjoint_smul_of_ne.trans <| forall_congr' fun _ ↦ or_iff_not_imp_left.symm + +@[to_additive] +lemma isBlock_iff_smul_eq_of_mem : + IsBlock G B ↔ ∀ ⦃g : G⦄ ⦃a : X⦄, a ∈ B → g • a ∈ B → g • B = B := by + simp [isBlock_iff_smul_eq_of_nonempty, Set.Nonempty, mem_smul_set] + +@[to_additive] alias ⟨IsBlock.disjoint_smul_of_ne, _⟩ := isBlock_iff_disjoint_smul_of_ne +@[to_additive] alias ⟨IsBlock.smul_eq_of_nonempty, _⟩ := isBlock_iff_smul_eq_of_nonempty +@[to_additive] alias ⟨IsBlock.smul_eq_or_disjoint, _⟩ := isBlock_iff_smul_eq_or_disjoint +@[to_additive] alias ⟨IsBlock.smul_eq_of_mem, _⟩ := isBlock_iff_smul_eq_of_mem + +-- TODO: Generalise to `SubgroupClass` +/-- If `B` is a `G`-block, then it is also a `H`-block for any subgroup `H` of `G`. -/ +@[to_additive +"If `B` is a `G`-block, then it is also a `H`-block for any subgroup `H` of `G`."] +lemma IsBlock.subgroup {H : Subgroup G} (hB : IsBlock G B) : IsBlock H B := fun _ _ h ↦ hB h + +/-- A block of a group action is invariant iff it is fixed. -/ +@[to_additive "A block of a group action is invariant iff it is fixed."] +lemma isInvariantBlock_iff_isFixedBlock : IsInvariantBlock G B ↔ IsFixedBlock G B := + ⟨fun hB g ↦ (hB g).antisymm <| subset_set_smul_iff.2 <| hB _, IsFixedBlock.isInvariantBlock⟩ + +/-- An invariant block of a group action is a fixed block. -/ +@[to_additive "An invariant block of a group action is a fixed block."] +alias ⟨IsInvariantBlock.isFixedBlock, _⟩ := isInvariantBlock_iff_isFixedBlock + +/-- An invariant block of a group action is a block. -/ +@[to_additive "An invariant block of a group action is a block."] +lemma IsInvariantBlock.isBlock (hB : IsInvariantBlock G B) : IsBlock G B := hB.isFixedBlock.isBlock + +/-- The full set is a fixed block. -/ +@[to_additive "The full set is a fixed block."] +lemma IsFixedBlock.univ : IsFixedBlock G (univ : Set X) := fun _ ↦ by simp + +/-- The full set is a block. -/ +@[to_additive (attr := simp) "The full set is a block."] +lemma IsBlock.univ : IsBlock G (univ : Set X) := IsFixedBlock.univ.isBlock + +/-- The intersection of two blocks is a block. -/ +@[to_additive "The intersection of two blocks is a block."] +lemma IsBlock.inter {B₁ B₂ : Set X} (h₁ : IsBlock G B₁) (h₂ : IsBlock G B₂) : + IsBlock G (B₁ ∩ B₂) := by + simp only [isBlock_iff_smul_eq_smul_of_nonempty, smul_set_inter] at h₁ h₂ ⊢ + rintro g₁ g₂ ⟨a, ha₁, ha₂⟩ + rw [h₁ ⟨a, ha₁.1, ha₂.1⟩, h₂ ⟨a, ha₁.2, ha₂.2⟩] -theorem IsBlock.smul_eq_or_disjoint {B : Set X} (hB : IsBlock G B) (g : G) : - g • B = B ∨ Disjoint (g • B) B := by - rw [IsBlock.def] at hB - simpa only [one_smul] using hB g 1 - -theorem IsBlock.def_one {B : Set X} : - IsBlock G B ↔ ∀ g : G, g • B = B ∨ Disjoint (g • B) B := by - refine ⟨IsBlock.smul_eq_or_disjoint, ?_⟩ - rw [IsBlock.def] - intro hB g g' - apply (hB (g'⁻¹ * g)).imp - · rw [← smul_smul, ← eq_inv_smul_iff, inv_inv] - exact id - · intro h - rw [Set.disjoint_iff] at h ⊢ - rintro x hx - suffices g'⁻¹ • x ∈ (g'⁻¹ * g) • B ∩ B by apply h this - rw [Set.mem_inter_iff] - simp only [← smul_smul, ← Set.mem_smul_set_iff_inv_smul_mem, smul_inv_smul] - exact hx - -theorem IsBlock.mk_notempty_one {B : Set X} : - IsBlock G B ↔ ∀ g : G, g • B ∩ B ≠ ∅ → g • B = B := by - simp_rw [IsBlock.def_one, Set.disjoint_iff_inter_eq_empty, or_iff_not_imp_right] - -theorem IsBlock.mk_mem {B : Set X} : - IsBlock G B ↔ ∀ (g : G) (a : X) (_ : a ∈ B) (_ : g • a ∈ B), g • B = B := by - rw [IsBlock.mk_notempty_one] - simp only [← Set.nonempty_iff_ne_empty, Set.nonempty_def, Set.mem_inter_iff, - exists_imp, and_imp, Set.mem_smul_set_iff_inv_smul_mem] - constructor - · intro H g a ha hga - apply H g (g • a) _ hga - simpa only [inv_smul_smul] using ha - · intro H g a ha hga - rw [← eq_inv_smul_iff, eq_comm] - exact H g⁻¹ a hga ha - -theorem IsBlock.def_mem {B : Set X} (hB : IsBlock G B) {a : X} {g : G} : - a ∈ B → g • a ∈ B → g • B = B := - IsBlock.mk_mem.mp hB g a - -theorem IsBlock.mk_subset {B : Set X} : - IsBlock G B ↔ ∀ {g : G} {b : X} (_ : b ∈ B) (_ : b ∈ g • B), g • B ⊆ B := by - simp_rw [IsBlock.mk_notempty_one, ← Set.nonempty_iff_ne_empty] - constructor - · intro hB g b hb hgb - exact (hB g ⟨b, hgb, hb⟩).le - · intro hB g ⟨b, hb', hb⟩ - apply le_antisymm (hB hb hb') - suffices g⁻¹ • B ≤ B by - rw [Set.le_iff_subset] at this ⊢ - rwa [← inv_inv g, ← Set.set_smul_subset_iff] - exact hB (Set.mem_smul_set_iff_inv_smul_mem.mp hb') (Set.smul_mem_smul_set_iff.mpr hb) - -/-- An invariant block is a fixed block -/ -theorem IsInvariantBlock.isFixedBlock {B : Set X} (hfB : IsInvariantBlock G B) : - IsFixedBlock G B := by - intro g - apply le_antisymm (hfB g) - intro x hx - rw [Set.mem_smul_set_iff_inv_smul_mem] - apply hfB g⁻¹ - rwa [Set.smul_mem_smul_set_iff] - -/-- An invariant block is a block -/ -theorem IsInvariantBlock.isBlock {B : Set X} (hfB : IsInvariantBlock G B) : - IsBlock G B := - hfB.isFixedBlock.isBlock - -/-- An orbit is a block -/ -theorem isFixedBlock_orbit (a : X) : IsFixedBlock G (orbit G a) := - (smul_orbit · a) - -/-- An orbit is a block -/ -theorem isBlock_orbit (a : X) : IsBlock G (orbit G a) := - (isFixedBlock_orbit a).isBlock - -variable (X) - -/-- The full set is a (trivial) block -/ -theorem isFixedBlock_univ : IsFixedBlock G (Set.univ : Set X) := - fun _ ↦ by simp only [Set.smul_set_univ] -@[deprecated (since := "2024-09-14")] alias isFixedBlock_top := isFixedBlock_univ - -/-- The full set is a (trivial) block -/ -theorem isBlock_univ : IsBlock G (Set.univ : Set X) := - (isFixedBlock_univ _).isBlock - -variable {X} - -/-- Is `B` is a block for an action of `G`, it is a block for the action of any subgroup of `G` -/ -theorem IsBlock.subgroup {H : Subgroup G} {B : Set X} (hfB : IsBlock G B) : - IsBlock H B := by - rw [IsBlock.def_one]; rintro ⟨g, _⟩ - simpa only using hfB.smul_eq_or_disjoint g - -theorem IsBlock.preimage {H Y : Type*} [Group H] [MulAction H Y] - {φ : H → G} (j : Y →ₑ[φ] X) {B : Set X} (hB : IsBlock G B) : +/-- An intersection of blocks is a block. -/ +@[to_additive "An intersection of blocks is a block."] +lemma IsBlock.iInter {ι : Sort*} {B : ι → Set X} (hB : ∀ i, IsBlock G (B i)) : + IsBlock G (⋂ i, B i) := by + simp only [isBlock_iff_smul_eq_smul_of_nonempty, smul_set_iInter] at hB ⊢ + rintro g₁ g₂ ⟨a, ha₁, ha₂⟩ + simp_rw [fun i ↦ hB i ⟨a, iInter_subset _ i ha₁, iInter_subset _ i ha₂⟩] + +/-- A trivial block is a block. -/ +@[to_additive "A trivial block is a block."] +lemma IsTrivialBlock.isBlock (hB : IsTrivialBlock B) : IsBlock G B := by + obtain hB | rfl := hB + · exact .of_subsingleton hB + · exact .univ + +/-- An orbit is a fixed block. -/ +@[to_additive "An orbit is a fixed block."] +protected lemma IsFixedBlock.orbit (a : X) : IsFixedBlock G (orbit G a) := (smul_orbit · a) + +/-- An orbit is a block. -/ +@[to_additive "An orbit is a block."] +protected lemma IsBlock.orbit (a : X) : IsBlock G (orbit G a) := (IsFixedBlock.orbit a).isBlock + +@[to_additive] +lemma isBlock_top : IsBlock (⊤ : Subgroup G) B ↔ IsBlock G B := + Subgroup.topEquiv.toEquiv.forall_congr fun _ ↦ Subgroup.topEquiv.toEquiv.forall_congr_left + +lemma IsBlock.preimage {H Y : Type*} [Group H] [MulAction H Y] + {φ : H → G} (j : Y →ₑ[φ] X) (hB : IsBlock G B) : IsBlock H (j ⁻¹' B) := by - rw [IsBlock.def_one] - intro g - rw [← Group.preimage_smul_setₛₗ Y X φ j] - apply (hB.smul_eq_or_disjoint (φ g)).imp - · intro heq - rw [heq] - · exact Disjoint.preimage _ + rintro g₁ g₂ hg + rw [← Group.preimage_smul_setₛₗ, ← Group.preimage_smul_setₛₗ] at hg ⊢ + exact (hB <| ne_of_apply_ne _ hg).preimage _ theorem IsBlock.image {H Y : Type*} [Group H] [MulAction H Y] {φ : G →* H} (j : X →ₑ[φ] Y) (hφ : Function.Surjective φ) (hj : Function.Injective j) - {B : Set X} (hB : IsBlock G B) : + (hB : IsBlock G B) : IsBlock H (j '' B) := by - rw [IsBlock.def] - intro h h' - obtain ⟨g, rfl⟩ := hφ h - obtain ⟨g', rfl⟩ := hφ h' - simp only [← image_smul_setₛₗ X Y φ j] - cases' IsBlock.def.mp hB g g' with h h - · left; rw [h] - · right; exact Set.disjoint_image_of_injective hj h - -theorem IsBlock.subtype_val_preimage {C : SubMulAction G X} {B : Set X} (hB : IsBlock G B) : + simp only [IsBlock, hφ.forall, ← image_smul_setₛₗ] + exact fun g₁ g₂ hg ↦ disjoint_image_of_injective hj <| hB <| ne_of_apply_ne _ hg + +theorem IsBlock.subtype_val_preimage {C : SubMulAction G X} (hB : IsBlock G B) : IsBlock G (Subtype.val ⁻¹' B : Set C) := hB.preimage C.inclusion -theorem IsBlock.iff_subtype_val {C : SubMulAction G X} {B : Set C} : - IsBlock G B ↔ IsBlock G (Subtype.val '' B : Set X) := by - simp only [IsBlock.def_one] - apply forall_congr' - intro g - rw [← SubMulAction.inclusion.coe_eq, ← image_smul_set _ _ _ C.inclusion g B, - ← Set.image_eq_image Subtype.coe_injective] - apply or_congr Iff.rfl - simp only [Set.disjoint_iff, Set.subset_empty_iff, Set.image_eq_empty, - ← C.inclusion_injective.injOn.image_inter (Set.subset_univ _) (Set.subset_univ _)] - -theorem IsBlock.iff_top (B : Set X) : - IsBlock G B ↔ IsBlock (⊤ : Subgroup G) B := by - simp only [IsBlock.def_one] - constructor - · intro h g; exact h g - · intro h g; exact h ⟨g, Subgroup.mem_top g⟩ +theorem isBlock_subtypeVal {C : SubMulAction G X} {B : Set C} : + IsBlock G (Subtype.val '' B : Set X) ↔ IsBlock G B := by + refine forall₂_congr fun g₁ g₂ ↦ ?_ + rw [← SubMulAction.inclusion.coe_eq, ← image_smul_set, ← image_smul_set, ne_eq, + Set.image_eq_image C.inclusion_injective, disjoint_image_iff C.inclusion_injective] -/-- The intersection of two blocks is a block -/ -theorem IsBlock.inter {B₁ B₂ : Set X} (h₁ : IsBlock G B₁) (h₂ : IsBlock G B₂) : - IsBlock G (B₁ ∩ B₂) := by - rw [IsBlock.def_one] - intro g - rw [Set.smul_set_inter] - cases' h₁.smul_eq_or_disjoint g with h₁ h₁ - · cases' h₂.smul_eq_or_disjoint g with h₂ h₂ - · left; rw [h₁, h₂] - right - apply Disjoint.inter_left'; apply Disjoint.inter_right' - exact h₂ - · right - apply Disjoint.inter_left; apply Disjoint.inter_right - exact h₁ - -/-- An intersection of blocks is a block -/ -theorem IsBlock.iInter {ι : Type*} {B : ι → Set X} (hB : ∀ i : ι, IsBlock G (B i)) : - IsBlock G (⋂ i, B i) := by - by_cases hι : (IsEmpty ι) - · -- ι = ∅, block = univ - suffices (⋂ i : ι, B i) = Set.univ by simpa only [this] using isBlock_univ X - simpa only [Set.iInter_eq_univ] using (hι.elim' ·) - rw [IsBlock.def_one] - intro g - rw [Set.smul_set_iInter] - by_cases h : ∃ i : ι, Disjoint (g • B i) (B i) - · right - obtain ⟨j, hj⟩ := h - refine Disjoint.mono ?_ ?_ hj <;> apply Set.iInter_subset - · left - simp only [not_exists] at h - have : ∀ i : ι, g • B i = B i := fun i => ((hB i).smul_eq_or_disjoint g).resolve_right (h i) - rw [Set.iInter_congr this] - -theorem IsBlock.of_subgroup_of_conjugate {B : Set X} {H : Subgroup G} (hB : IsBlock H B) (g : G) : +theorem IsBlock.of_subgroup_of_conjugate {H : Subgroup G} (hB : IsBlock H B) (g : G) : IsBlock (Subgroup.map (MulEquiv.toMonoidHom (MulAut.conj g)) H) (g • B) := by - rw [IsBlock.def_one] + rw [isBlock_iff_smul_eq_or_disjoint] intro h' obtain ⟨h, hH, hh⟩ := Subgroup.mem_map.mp (SetLike.coe_mem h') simp only [MulEquiv.coe_toMonoidHom, MulAut.conj_apply] at hh @@ -329,9 +283,9 @@ theorem IsBlock.of_subgroup_of_conjugate {B : Set X} {H : Subgroup G} (hB : IsBl rw [← hh, smul_smul (g * h * g⁻¹) g B, smul_smul g h B, inv_mul_cancel_right] /-- A translate of a block is a block -/ -theorem IsBlock.translate {B : Set X} (g : G) (hB : IsBlock G B) : +theorem IsBlock.translate (g : G) (hB : IsBlock G B) : IsBlock G (g • B) := by - rw [IsBlock.iff_top] at hB ⊢ + rw [← isBlock_top] at hB ⊢ rw [← Subgroup.map_comap_eq_self_of_surjective (f := MulAut.conj g) (MulAut.conj g).surjective ⊤] apply IsBlock.of_subgroup_of_conjugate rwa [Subgroup.comap_top] @@ -339,12 +293,11 @@ theorem IsBlock.translate {B : Set X} (g : G) (hB : IsBlock G B) : variable (G) in /-- For `SMul G X`, a block system of `X` is a partition of `X` into blocks for the action of `G` -/ -def IsBlockSystem (B : Set (Set X)) := - Setoid.IsPartition B ∧ ∀ b : Set X, b ∈ B → IsBlock G b +def IsBlockSystem (ℬ : Set (Set X)) := Setoid.IsPartition ℬ ∧ ∀ ⦃B⦄, B ∈ ℬ → IsBlock G B -/-- Translates of a block form a `block_system` -/ +/-- Translates of a block form a block system. -/ theorem IsBlock.isBlockSystem [hGX : MulAction.IsPretransitive G X] - {B : Set X} (hB : IsBlock G B) (hBe : B.Nonempty) : + (hB : IsBlock G B) (hBe : B.Nonempty) : IsBlockSystem G (Set.range fun g : G => g • B) := by refine ⟨⟨?nonempty, ?cover⟩, ?mem_blocks⟩ case mem_blocks => rintro B' ⟨g, rfl⟩; exact hB.translate g @@ -359,10 +312,7 @@ theorem IsBlock.isBlockSystem [hGX : MulAction.IsPretransitive G X] simp only [Set.smul_mem_smul_set_iff, hb, exists_unique_iff_exists, Set.mem_range, exists_apply_eq_apply, exists_const, exists_prop, and_imp, forall_exists_index, forall_apply_eq_imp_iff, true_and] - intro g' ha - apply (IsBlock.def.mp hB g' g).resolve_right - rw [Set.not_disjoint_iff] - refine ⟨g • b, ha, ⟨b, hb, rfl⟩⟩ + exact fun g' ha ↦ hB.smul_eq_smul_of_nonempty ⟨g • b, ha, ⟨b, hb, rfl⟩⟩ section Normal @@ -382,9 +332,9 @@ lemma smul_orbit_eq_orbit_smul (N : Subgroup G) [nN : N.Normal] (a : X) (g : G) simp only [← mul_assoc, ← smul_smul, smul_inv_smul, inv_inv] /-- An orbit of a normal subgroup is a block -/ -theorem orbit.isBlock_of_normal {N : Subgroup G} [N.Normal] (a : X) : +theorem IsBlock.orbit_of_normal {N : Subgroup G} [N.Normal] (a : X) : IsBlock G (orbit N a) := by - rw [IsBlock.def_one] + rw [isBlock_iff_smul_eq_or_disjoint] intro g rw [smul_orbit_eq_orbit_smul] apply orbit.eq_or_disjoint @@ -395,7 +345,48 @@ theorem IsBlockSystem.of_normal {N : Subgroup G} [N.Normal] : constructor · apply IsPartition.of_orbits · intro b; rintro ⟨a, rfl⟩ - exact orbit.isBlock_of_normal a + exact .orbit_of_normal a + +section Group +variable {S H : Type*} [Group H] [SetLike S H] [SubgroupClass S H] {s : S} {a b : G} + +/-! +Annoyingly, it seems like the following two lemmas cannot be unified. +-/ + +section Left +variable [MulAction G H] [IsScalarTower G H H] + +/-- See `MulAction.isBlock_subgroup'` for a version that works for the right action of a group on +itself. -/ +@[to_additive "See `AddAction.isBlock_subgroup'` for a version that works for the right action +of a group on itself."] +lemma isBlock_subgroup : IsBlock G (s : Set H) := by + simp only [IsBlock, disjoint_left] + rintro a b hab _ ⟨c, hc, rfl⟩ ⟨d, hd, (hcd : b • d = a • c)⟩ + refine hab ?_ + rw [← smul_coe_set hc, ← smul_assoc, ← hcd, smul_assoc, smul_coe_set hc, smul_coe_set hd] + +end Left + +section Right +variable [MulAction G H] [IsScalarTower G Hᵐᵒᵖ H] + +open MulOpposite + +/-- See `MulAction.isBlock_subgroup` for a version that works for the left action of a group on +itself. -/ +@[to_additive "See `AddAction.isBlock_subgroup` for a version that works for the left action +of a group on itself."] +lemma isBlock_subgroup' : IsBlock G (s : Set H) := by + simp only [IsBlock, disjoint_left] + rintro a b hab _ ⟨c, hc, rfl⟩ ⟨d, hd, (hcd : b • d = a • c)⟩ + refine hab ?_ + rw [← op_smul_coe_set hc, ← smul_assoc, ← op_smul, ← hcd, op_smul, smul_assoc, op_smul_coe_set hc, + op_smul_coe_set hd] + +end Right +end Group end Normal @@ -411,28 +402,21 @@ section Stabilizer /-- The orbit of `a` under a subgroup containing the stabilizer of `a` is a block -/ theorem IsBlock.of_orbit {H : Subgroup G} {a : X} (hH : stabilizer G a ≤ H) : IsBlock G (MulAction.orbit H a) := by - simp_rw [IsBlock.def_one, or_iff_not_imp_right, Set.not_disjoint_iff] - rintro g ⟨-, ⟨-, ⟨h₁, rfl⟩, h⟩, ⟨h₂, rfl⟩⟩ + rw [isBlock_iff_smul_eq_of_nonempty] + rintro g ⟨-, ⟨-, ⟨h₁, rfl⟩, h⟩, h₂, rfl⟩ suffices g ∈ H by rw [← Subgroup.coe_mk H g this, ← H.toSubmonoid.smul_def, smul_orbit (⟨g, this⟩ : H) a] rw [← mul_mem_cancel_left h₂⁻¹.2, ← mul_mem_cancel_right h₁.2] apply hH - simp only [mem_stabilizer_iff, InvMemClass.coe_inv, mul_smul, inv_smul_eq_iff] - exact h + simpa only [mem_stabilizer_iff, InvMemClass.coe_inv, mul_smul, inv_smul_eq_iff] /-- If `B` is a block containing `a`, then the stabilizer of `B` contains the stabilizer of `a` -/ -theorem IsBlock.stabilizer_le {B : Set X} (hB : IsBlock G B) {a : X} (ha : a ∈ B) : - stabilizer G a ≤ stabilizer G B := by - intro g hg - apply Or.resolve_right (hB.smul_eq_or_disjoint g) - rw [Set.not_disjoint_iff] - refine ⟨a, ?_, ha⟩ - rw [← hg, Set.smul_mem_smul_set_iff] - exact ha +theorem IsBlock.stabilizer_le (hB : IsBlock G B) {a : X} (ha : a ∈ B) : + stabilizer G a ≤ stabilizer G B := + fun g hg ↦ hB.smul_eq_of_nonempty ⟨a, by rwa [← hg, smul_mem_smul_set_iff], ha⟩ /-- A block containing `a` is the orbit of `a` under its stabilizer -/ -theorem IsBlock.orbit_stabilizer_eq - [htGX : IsPretransitive G X] {B : Set X} (hB : IsBlock G B) {a : X} (ha : a ∈ B) : +theorem IsBlock.orbit_stabilizer_eq [IsPretransitive G X] (hB : IsBlock G B) {a : X} (ha : a ∈ B) : MulAction.orbit (stabilizer G B) a = B := by ext x constructor @@ -442,7 +426,7 @@ theorem IsBlock.orbit_stabilizer_eq exact ha · intro hx obtain ⟨k, rfl⟩ := exists_smul_eq G a x - exact ⟨⟨k, hB.def_mem ha hx⟩, rfl⟩ + exact ⟨⟨k, hB.smul_eq_of_mem ha hx⟩, rfl⟩ /-- A subgroup containing the stabilizer of `a` is the stabilizer of the orbit of `a` under that subgroup -/ @@ -479,9 +463,9 @@ def block_stabilizerOrderIso [htGX : IsPretransitive G X] (a : X) : obtain ⟨k, rfl⟩ := htGX.exists_smul_eq a b suffices k ∈ stabilizer G B' by exact this.symm ▸ (Set.smul_mem_smul_set ha') - exact hBB' (hB.def_mem ha hb) + exact hBB' (hB.smul_eq_of_mem ha hb) · intro hBB' g hgB - apply IsBlock.def_mem hB' ha' + apply hB'.smul_eq_of_mem ha' exact hBB' <| hgB.symm ▸ (Set.smul_mem_smul_set ha) end Stabilizer @@ -552,7 +536,7 @@ theorem of_subset {B : Set X} (a : X) (hfB : B.Finite) : IsBlock G (⋂ (k : G) (_ : a ∈ k • B), k • B) := by let B' := ⋂ (k : G) (_ : a ∈ k • B), k • B cases' Set.eq_empty_or_nonempty B with hfB_e hfB_ne - · simp [hfB_e, isBlock_univ] + · simp [hfB_e] have hB'₀ : ∀ (k : G) (_ : a ∈ k • B), B' ≤ k • B := by intro k hk exact Set.biInter_subset_of_mem hk @@ -569,10 +553,8 @@ theorem of_subset {B : Set X} (a : X) (hfB : B.Finite) : have hag' (g : G) (hg : a ∈ g • B') : B' = g • B' := by rw [eq_comm, ← mem_stabilizer_iff, mem_stabilizer_of_finite_iff_le_smul _ hfB'] exact hag g hg - rw [mk_notempty_one] - intro g hg - rw [← Set.nonempty_iff_ne_empty] at hg - obtain ⟨b : X, hb' : b ∈ g • B', hb : b ∈ B'⟩ := Set.nonempty_def.mp hg + rw [isBlock_iff_smul_eq_of_nonempty] + rintro g ⟨b : X, hb' : b ∈ g • B', hb : b ∈ B'⟩ obtain ⟨k : G, hk : k • a = b⟩ := exists_smul_eq G a b have hak : a ∈ k⁻¹ • B' := by refine ⟨b, hb, ?_⟩ From 87bea92ed780da877a30bad831855d40a9dad540 Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Tue, 22 Oct 2024 16:44:11 +0000 Subject: [PATCH 285/505] feat(NumberTheory/LSeries): functional equation for Dirichlet L-funcs (#16967) Define completed L-functions for Dirichlet characters, prove that they are differentiable (except when they obviously aren't), and prove the functional equation. --- Mathlib/Analysis/Fourier/ZMod.lean | 7 +- .../DirichletCharacter/Basic.lean | 25 ++- .../LSeries/DirichletContinuation.lean | 160 +++++++++++++++++- Mathlib/NumberTheory/LSeries/ZMod.lean | 5 + 4 files changed, 188 insertions(+), 9 deletions(-) diff --git a/Mathlib/Analysis/Fourier/ZMod.lean b/Mathlib/Analysis/Fourier/ZMod.lean index 4fad478f0bc2c..ae3c4f996bbcc 100644 --- a/Mathlib/Analysis/Fourier/ZMod.lean +++ b/Mathlib/Analysis/Fourier/ZMod.lean @@ -201,9 +201,9 @@ end ZMod namespace DirichletCharacter -variable {N : ℕ} [NeZero N] (χ : DirichletCharacter ℂ N) +variable {N : ℕ} [NeZero N] -lemma fourierTransform_eq_gaussSum_mulShift (k : ZMod N) : +lemma fourierTransform_eq_gaussSum_mulShift (χ : DirichletCharacter ℂ N) (k : ZMod N) : 𝓕 χ k = gaussSum χ (stdAddChar.mulShift (-k)) := by simp only [dft_apply, smul_eq_mul] congr 1 with j @@ -211,7 +211,8 @@ lemma fourierTransform_eq_gaussSum_mulShift (k : ZMod N) : /-- For a primitive Dirichlet character `χ`, the Fourier transform of `χ` is a constant multiple of `χ⁻¹` (and the constant is essentially the Gauss sum). -/ -lemma fourierTransform_eq_inv_mul_gaussSum (k : ZMod N) (hχ : IsPrimitive χ) : +lemma IsPrimitive.fourierTransform_eq_inv_mul_gaussSum {χ : DirichletCharacter ℂ N} + (hχ : IsPrimitive χ) (k : ZMod N) : 𝓕 χ k = χ⁻¹ (-k) * gaussSum χ stdAddChar := by rw [fourierTransform_eq_gaussSum_mulShift, gaussSum_mulShift_of_isPrimitive _ hχ] diff --git a/Mathlib/NumberTheory/DirichletCharacter/Basic.lean b/Mathlib/NumberTheory/DirichletCharacter/Basic.lean index 9fbae12e56bcd..8780485d9254f 100644 --- a/Mathlib/NumberTheory/DirichletCharacter/Basic.lean +++ b/Mathlib/NumberTheory/DirichletCharacter/Basic.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Ashvni Narayanan. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Ashvni Narayanan, Moritz Firsching, Michael Stoll -/ +import Mathlib.Algebra.Group.EvenFunction import Mathlib.Data.ZMod.Units import Mathlib.NumberTheory.MulChar.Basic @@ -159,6 +160,10 @@ instance : Subsingleton (DirichletCharacter R 1) := by noncomputable instance : Unique (DirichletCharacter R 1) := Unique.mk' (DirichletCharacter R 1) +/-- A Dirichlet character of modulus `≠ 1` maps `0` to `0`. -/ +lemma map_zero' (hn : n ≠ 1) : χ 0 = 0 := + have := ZMod.nontrivial_iff.mpr hn; χ.map_zero + lemma changeLevel_one {d : ℕ} (h : d ∣ n) : changeLevel h (1 : DirichletCharacter R d) = 1 := by simp @@ -287,7 +292,7 @@ lemma primitive_mul_isPrimitive {m : ℕ} (ψ : DirichletCharacter R m) : section CommRing -variable {S : Type} [CommRing S] {m : ℕ} (ψ : DirichletCharacter S m) +variable {S : Type*} [CommRing S] {m : ℕ} (ψ : DirichletCharacter S m) /-- A Dirichlet character is odd if its value at -1 is -1. -/ def Odd : Prop := ψ (-1) = -1 @@ -299,6 +304,16 @@ lemma even_or_odd [NoZeroDivisors S] : ψ.Even ∨ ψ.Odd := by suffices ψ (-1) ^ 2 = 1 by convert sq_eq_one_iff.mp this rw [← map_pow _, neg_one_sq, map_one] +lemma not_even_and_odd [NeZero (2 : S)] : ¬(ψ.Even ∧ ψ.Odd) := by + rintro ⟨(h : _ = 1), (h' : _ = -1)⟩ + simp only [h', neg_eq_iff_add_eq_zero, one_add_one_eq_two, two_ne_zero] at h + +lemma Even.not_odd [NeZero (2 : S)] (hψ : Even ψ) : ¬Odd ψ := + not_and.mp ψ.not_even_and_odd hψ + +lemma Odd.not_even [NeZero (2 : S)] (hψ : Odd ψ) : ¬Even ψ := + not_and'.mp ψ.not_even_and_odd hψ + lemma Odd.toUnitHom_eval_neg_one (hψ : ψ.Odd) : ψ.toUnitHom (-1) = -1 := by rw [← Units.eq_iff, MulChar.coe_toUnitHom] exact hψ @@ -317,6 +332,14 @@ lemma Even.eval_neg (x : ZMod m) (hψ : ψ.Even) : ψ (- x) = ψ x := by rw [← neg_one_mul, map_mul] simp [hψ] +/-- An even Dirichlet character is an even function. -/ +lemma Even.to_fun {χ : DirichletCharacter S m} (hχ : Even χ) : Function.Even χ := + fun _ ↦ by rw [← neg_one_mul, map_mul, hχ, one_mul] + +/-- An odd Dirichlet character is an odd function. -/ +lemma Odd.to_fun {χ : DirichletCharacter S m} (hχ : Odd χ) : Function.Odd χ := + fun _ ↦ by rw [← neg_one_mul, map_mul, hχ, neg_one_mul] + end CommRing end DirichletCharacter diff --git a/Mathlib/NumberTheory/LSeries/DirichletContinuation.lean b/Mathlib/NumberTheory/LSeries/DirichletContinuation.lean index dd4b132bc5130..2601a88444d1f 100644 --- a/Mathlib/NumberTheory/LSeries/DirichletContinuation.lean +++ b/Mathlib/NumberTheory/LSeries/DirichletContinuation.lean @@ -10,22 +10,36 @@ import Mathlib.NumberTheory.DirichletCharacter.Basic # Analytic continuation of Dirichlet L-functions We show that if `χ` is a Dirichlet character `ZMod N → ℂ`, for a positive integer `N`, then the -L-series of `χ` has analytic continuation (away from a pole at `s = 1` if `χ` is trivial). +L-series of `χ` has analytic continuation (away from a pole at `s = 1` if `χ` is trivial), and +similarly for completed L-functions. All definitions and theorems are in the `DirichletCharacter` namespace. ## Main definitions * `LFunction χ s`: the L-function, defined as a linear combination of Hurwitz zeta functions. +* `completedLFunction χ s`: the completed L-function, which for *almost* all `s` is equal to + `LFunction χ s * gammaFactor χ s` where `gammaFactor χ s` is the archimedean Gamma-factor. +* `rootNumber`: the global root number of the L-series of `χ` (for `χ` primitive; junk otherwise). ## Main theorems * `LFunction_eq_LSeries`: if `1 < re s` then the `LFunction` coincides with the naive `LSeries`. * `differentiable_LFunction`: if `χ` is nontrivial then `LFunction χ s` is differentiable everywhere. +* `LFunction_eq_completed_div_gammaFactor`: we have + `LFunction χ s = completedLFunction χ s / gammaFactor χ s`, unless `s = 0` and `χ` is the trivial + character modulo 1. +* `differentiable_completedLFunction`: if `χ` is nontrivial then `completedLFunction χ s` is + differentiable everywhere. +* `IsPrimitive.completedLFunction_one_sub`: the **functional equation** for Dirichlet L-functions, + showing that if `χ` is primitive modulo `N`, then + `completedLFunction χ s = N ^ (s - 1 / 2) * rootNumber χ * completedLFunction χ⁻¹ s`. -/ -open Complex +open HurwitzZeta Complex Finset Classical ZMod + +open scoped Real namespace DirichletCharacter @@ -38,13 +52,16 @@ latter is convergent. This is constructed as a linear combination of Hurwitz zet Note that this is not the same as `LSeries χ`: they agree in the convergence range, but `LSeries χ s` is defined to be `0` if `re s ≤ 1`. -/ +@[pp_nodot] noncomputable def LFunction (χ : DirichletCharacter ℂ N) (s : ℂ) : ℂ := ZMod.LFunction χ s -/-- The L-function of the (unique) Dirichlet character mod 1 is the Riemann zeta function. -(Compare `DirichletCharacter.LSeries_modOne_eq`.) -/ +/-- +The L-function of the (unique) Dirichlet character mod 1 is the Riemann zeta function. +(Compare `DirichletCharacter.LSeries_modOne_eq`.) +-/ @[simp] lemma LFunction_modOne_eq {χ : DirichletCharacter ℂ 1} : LFunction χ = riemannZeta := by - ext1; rw [LFunction, ZMod.LFunction_modOne_eq, (by rfl : (0 : ZMod 1) = 1), map_one, one_mul] + ext; rw [LFunction, ZMod.LFunction_modOne_eq, (by rfl : (0 : ZMod 1) = 1), map_one, one_mul] /-- For `1 < re s` the L-function of a Dirichlet character agrees with the sum of the naive Dirichlet @@ -67,4 +84,137 @@ lemma differentiable_LFunction {χ : DirichletCharacter ℂ N} (hχ : χ ≠ 1) Differentiable ℂ (LFunction χ) := (differentiableAt_LFunction _ · <| Or.inr hχ) +/-- The L-function of an even Dirichlet character vanishes at strictly negative even integers. -/ +@[simp] +lemma Even.LFunction_neg_two_mul_nat_add_one {χ : DirichletCharacter ℂ N} (hχ : Even χ) (n : ℕ) : + LFunction χ (-(2 * (n + 1))) = 0 := + ZMod.LFunction_neg_two_mul_nat_add_one hχ.to_fun n + +/-- The L-function of an even Dirichlet character vanishes at strictly negative even integers. -/ +@[simp] +lemma Even.LFunction_neg_two_mul_nat {χ : DirichletCharacter ℂ N} (hχ : Even χ) (n : ℕ) [NeZero n] : + LFunction χ (-(2 * n)) = 0 := by + obtain ⟨m, rfl⟩ := Nat.exists_eq_succ_of_ne_zero (NeZero.ne n) + exact_mod_cast hχ.LFunction_neg_two_mul_nat_add_one m + +/-- The L-function of an odd Dirichlet character vanishes at negative odd integers. -/ +@[simp] lemma Odd.LFunction_neg_two_mul_nat_sub_one + {χ : DirichletCharacter ℂ N} (hχ : Odd χ) (n : ℕ) : + LFunction χ (-(2 * n) - 1) = 0 := + ZMod.LFunction_neg_two_mul_nat_sub_one hχ.to_fun n + +section gammaFactor + +omit [NeZero N] -- not required for these declarations + +/-- The Archimedean Gamma factor: `Gammaℝ s` if `χ` is even, and `Gammaℝ (s + 1)` otherwise. -/ +noncomputable def gammaFactor (χ : DirichletCharacter ℂ N) (s : ℂ) := + if χ.Even then Gammaℝ s else Gammaℝ (s + 1) + +lemma Even.gammaFactor_def {χ : DirichletCharacter ℂ N} (hχ : χ.Even) (s : ℂ) : + gammaFactor χ s = Gammaℝ s := by + simp only [gammaFactor, hχ, ↓reduceIte] + +lemma Odd.gammaFactor_def {χ : DirichletCharacter ℂ N} (hχ : χ.Odd) (s : ℂ) : + gammaFactor χ s = Gammaℝ (s + 1) := by + simp only [gammaFactor, hχ.not_even, ↓reduceIte] + +end gammaFactor + +/-- +The completed L-function of a Dirichlet character, almost everywhere equal to +`LFunction χ s * gammaFactor χ s`. +-/ +@[pp_nodot] noncomputable def completedLFunction (χ : DirichletCharacter ℂ N) (s : ℂ) : ℂ := + ZMod.completedLFunction χ s + +/-- +The completed L-function of the (unique) Dirichlet character mod 1 is the completed Riemann zeta +function. +-/ +lemma completedLFunction_modOne_eq {χ : DirichletCharacter ℂ 1} : + completedLFunction χ = completedRiemannZeta := by + ext; rw [completedLFunction, ZMod.completedLFunction_modOne_eq, map_one, one_mul] + +/-- +The completed L-function of a Dirichlet character is differentiable, with the following +exceptions: at `s = 1` if `χ` is the trivial character (to any modulus); and at `s = 0` if the +modulus is 1. This result is best possible. + +Note both `χ` and `s` are explicit arguments: we will always be able to infer one or other +of them from the hypotheses, but it's not clear which! +-/ +lemma differentiableAt_completedLFunction (χ : DirichletCharacter ℂ N) (s : ℂ) + (hs₀ : s ≠ 0 ∨ N ≠ 1) (hs₁ : s ≠ 1 ∨ χ ≠ 1) : + DifferentiableAt ℂ (completedLFunction χ) s := + ZMod.differentiableAt_completedLFunction _ _ (by have := χ.map_zero'; tauto) + (by have := χ.sum_eq_zero_of_ne_one; tauto) + +/-- The completed L-function of a non-trivial Dirichlet character is differentiable everywhere. -/ +lemma differentiable_completedLFunction {χ : DirichletCharacter ℂ N} (hχ : χ ≠ 1) : + Differentiable ℂ (completedLFunction χ) := by + refine fun s ↦ differentiableAt_completedLFunction _ _ (Or.inr ?_) (Or.inr hχ) + exact hχ ∘ level_one' _ + +/-- +Relation between the completed L-function and the usual one. We state it this way around so +it holds at the poles of the gamma factor as well. +-/ +lemma LFunction_eq_completed_div_gammaFactor (χ : DirichletCharacter ℂ N) (s : ℂ) + (h : s ≠ 0 ∨ N ≠ 1) : LFunction χ s = completedLFunction χ s / gammaFactor χ s := by + rcases χ.even_or_odd with hχ | hχ <;> + rw [hχ.gammaFactor_def] + · exact LFunction_eq_completed_div_gammaFactor_even hχ.to_fun _ (h.imp_right χ.map_zero') + · apply LFunction_eq_completed_div_gammaFactor_odd hχ.to_fun + +/-- +Global root number of `χ` (for `χ` primitive; junk otherwise). Defined as +`gaussSum χ stdAddChar / I ^ a / N ^ (1 / 2)`, where `a = 0` if even, `a = 1` if odd. (The factor +`1 / I ^ a` is the Archimedean root number.) This is a complex number of absolute value 1. +-/ +noncomputable def rootNumber (χ : DirichletCharacter ℂ N) : ℂ := + gaussSum χ stdAddChar / I ^ (if χ.Even then 0 else 1) / N ^ (1 / 2 : ℂ) + +/-- The root number of the unique Dirichlet character modulo 1 is 1. -/ +lemma rootNumber_modOne (χ : DirichletCharacter ℂ 1) : rootNumber χ = 1 := by + simp only [rootNumber, gaussSum, ← singleton_eq_univ (1 : ZMod 1), sum_singleton, map_one, + (show stdAddChar (1 : ZMod 1) = 1 from AddChar.map_zero_eq_one _), one_mul, + (show χ.Even from map_one _), ite_true, pow_zero, div_one, Nat.cast_one, one_cpow] + +namespace IsPrimitive + +/-- **Functional equation** for primitive Dirichlet L-functions. -/ +theorem completedLFunction_one_sub {χ : DirichletCharacter ℂ N} (hχ : IsPrimitive χ) (s : ℂ) : + completedLFunction χ (1 - s) = N ^ (s - 1 / 2) * rootNumber χ * completedLFunction χ⁻¹ s := by + -- First handle special case of Riemann zeta + rcases eq_or_ne N 1 with rfl | hN + · simp only [completedLFunction_modOne_eq, completedRiemannZeta_one_sub, Nat.cast_one, one_cpow, + rootNumber_modOne, one_mul] + -- facts about `χ` as function + have h_sum : ∑ j, χ j = 0 := by + refine χ.sum_eq_zero_of_ne_one (fun h ↦ hN.symm ?_) + rwa [IsPrimitive, h, conductor_one (NeZero.ne _)] at hχ + let ε := I ^ (if χ.Even then 0 else 1) + -- gather up powers of N + rw [rootNumber, ← mul_comm_div, ← mul_comm_div, ← cpow_sub _ _ (NeZero.ne _), sub_sub, add_halves] + calc completedLFunction χ (1 - s) + _ = N ^ (s - 1) * χ (-1) / ε * ZMod.completedLFunction (𝓕 χ) s := by + simp only [ε] + split_ifs with h + · rw [pow_zero, div_one, h, mul_one, completedLFunction, + completedLFunction_one_sub_even h.to_fun _ (.inr h_sum) (.inr <| χ.map_zero' hN)] + · replace h : χ.Odd := χ.even_or_odd.resolve_left h + rw [completedLFunction, completedLFunction_one_sub_odd h.to_fun, + pow_one, h, div_I, mul_neg_one, ← neg_mul, neg_neg] + _ = (_) * ZMod.completedLFunction (fun j ↦ χ⁻¹ (-1) * gaussSum χ stdAddChar * χ⁻¹ j) s := by + congr 2 with j + rw [hχ.fourierTransform_eq_inv_mul_gaussSum, ← neg_one_mul j, map_mul, mul_right_comm] + _ = N ^ (s - 1) / ε * gaussSum χ stdAddChar * completedLFunction χ⁻¹ s * (χ (-1) * χ⁻¹ (-1)):= by + rw [completedLFunction, completedLFunction_const_mul] + ring + _ = N ^ (s - 1) / ε * gaussSum χ stdAddChar * completedLFunction χ⁻¹ s := by + rw [← MulChar.mul_apply, mul_inv_cancel, MulChar.one_apply (isUnit_one.neg), mul_one] + +end IsPrimitive + end DirichletCharacter diff --git a/Mathlib/NumberTheory/LSeries/ZMod.lean b/Mathlib/NumberTheory/LSeries/ZMod.lean index 2ef6c910b132d..9535bbf835aa8 100644 --- a/Mathlib/NumberTheory/LSeries/ZMod.lean +++ b/Mathlib/NumberTheory/LSeries/ZMod.lean @@ -294,6 +294,11 @@ noncomputable def completedLFunction (Φ : ZMod N → ℂ) (s : ℂ) : ℂ := @[simp] lemma completedLFunction_zero (s : ℂ) : completedLFunction (0 : ZMod N → ℂ) s = 0 := by simp only [completedLFunction, Pi.zero_apply, zero_mul, sum_const_zero, mul_zero, zero_add] +lemma completedLFunction_const_mul (a : ℂ) (Φ : ZMod N → ℂ) (s : ℂ) : + completedLFunction (fun j ↦ a * Φ j) s = a * completedLFunction Φ s := by + simp only [completedLFunction, mul_add, mul_sum] + congr with i <;> ring + lemma completedLFunction_def_even (hΦ : Φ.Even) (s : ℂ) : completedLFunction Φ s = N ^ (-s) * ∑ j, Φ j * completedHurwitzZetaEven (toAddCircle j) s := by suffices ∑ j, Φ j * completedHurwitzZetaOdd (toAddCircle j) s = 0 by From 93c444d301e86a998f4e8247dde5d8977258b108 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Tue, 22 Oct 2024 16:44:13 +0000 Subject: [PATCH 286/505] chore(Data/ENNReal): split off `Lemmas` file (#17919) Almost all of `Data.ENNReal.Basic` can be done with much a smaller set of imports. We only have two lemmas that need more imports, so let's split them off to a new file. Co-authored-by: Anne Baanen --- Mathlib.lean | 1 + Mathlib/Data/ENNReal/Basic.lean | 12 ------ Mathlib/Data/ENNReal/Lemmas.lean | 37 +++++++++++++++++++ Mathlib/Topology/MetricSpace/Pseudo/Pi.lean | 1 + .../MetricSpace/ThickenedIndicator.lean | 2 +- 5 files changed, 40 insertions(+), 13 deletions(-) create mode 100644 Mathlib/Data/ENNReal/Lemmas.lean diff --git a/Mathlib.lean b/Mathlib.lean index 0f583235c97b4..fe64af1dced4a 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2192,6 +2192,7 @@ import Mathlib.Data.DFinsupp.WellFounded import Mathlib.Data.DList.Instances import Mathlib.Data.ENNReal.Basic import Mathlib.Data.ENNReal.Inv +import Mathlib.Data.ENNReal.Lemmas import Mathlib.Data.ENNReal.Operations import Mathlib.Data.ENNReal.Real import Mathlib.Data.ENat.Basic diff --git a/Mathlib/Data/ENNReal/Basic.lean b/Mathlib/Data/ENNReal/Basic.lean index 0f906e0ef21a5..737f546515208 100644 --- a/Mathlib/Data/ENNReal/Basic.lean +++ b/Mathlib/Data/ENNReal/Basic.lean @@ -3,12 +3,9 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Yury Kudryashov -/ -import Mathlib.Algebra.Group.Indicator import Mathlib.Algebra.Order.Ring.WithTop import Mathlib.Algebra.Order.Sub.WithTop -import Mathlib.Data.Finset.Lattice import Mathlib.Data.NNReal.Defs -import Mathlib.Order.Interval.Set.WithBotTop /-! # Extended non-negative reals @@ -438,11 +435,6 @@ def ofNNRealHom : ℝ≥0 →+* ℝ≥0∞ where @[simp] theorem coe_ofNNRealHom : ⇑ofNNRealHom = some := rfl -@[simp, norm_cast] -theorem coe_indicator {α} (s : Set α) (f : α → ℝ≥0) (a : α) : - ((s.indicator f a : ℝ≥0) : ℝ≥0∞) = s.indicator (fun x => ↑(f x)) a := - (ofNNRealHom : ℝ≥0 →+ ℝ≥0∞).map_indicator _ _ _ - section Order theorem bot_eq_zero : (⊥ : ℝ≥0∞) = 0 := rfl @@ -504,10 +496,6 @@ theorem toReal_le_coe_of_le_coe {a : ℝ≥0∞} {b : ℝ≥0} (h : a ≤ b) : a lift a to ℝ≥0 using ne_top_of_le_ne_top coe_ne_top h simpa using h -@[simp, norm_cast] -theorem coe_finset_sup {s : Finset α} {f : α → ℝ≥0} : ↑(s.sup f) = s.sup fun x => (f x : ℝ≥0∞) := - Finset.comp_sup_eq_sup_comp_of_is_total _ coe_mono rfl - @[simp] theorem max_eq_zero_iff : max a b = 0 ↔ a = 0 ∧ b = 0 := max_eq_bot theorem max_zero_left : max 0 a = a := diff --git a/Mathlib/Data/ENNReal/Lemmas.lean b/Mathlib/Data/ENNReal/Lemmas.lean new file mode 100644 index 0000000000000..979bf8305348b --- /dev/null +++ b/Mathlib/Data/ENNReal/Lemmas.lean @@ -0,0 +1,37 @@ +/- +Copyright (c) 2017 Johannes Hölzl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johannes Hölzl, Yury Kudryashov +-/ +import Mathlib.Algebra.Group.Indicator +import Mathlib.Data.ENNReal.Basic +import Mathlib.Data.Finset.Lattice + +/-! +# Some lemmas on extended non-negative reals + +These are some lemmas split off from `ENNReal.Basic` because they need a lot more imports. +They are probably good targets for further cleanup or moves. +-/ + + +open Function Set NNReal + +variable {α : Type*} + +namespace ENNReal + +@[simp, norm_cast] +theorem coe_indicator {α} (s : Set α) (f : α → ℝ≥0) (a : α) : + ((s.indicator f a : ℝ≥0) : ℝ≥0∞) = s.indicator (fun x => ↑(f x)) a := + (ofNNRealHom : ℝ≥0 →+ ℝ≥0∞).map_indicator _ _ _ + +section Order + +@[simp, norm_cast] +theorem coe_finset_sup {s : Finset α} {f : α → ℝ≥0} : ↑(s.sup f) = s.sup fun x => (f x : ℝ≥0∞) := + Finset.comp_sup_eq_sup_comp_of_is_total _ coe_mono rfl + +end Order + +end ENNReal diff --git a/Mathlib/Topology/MetricSpace/Pseudo/Pi.lean b/Mathlib/Topology/MetricSpace/Pseudo/Pi.lean index 92a7c3925b2a1..c7b60bab973a9 100644 --- a/Mathlib/Topology/MetricSpace/Pseudo/Pi.lean +++ b/Mathlib/Topology/MetricSpace/Pseudo/Pi.lean @@ -3,6 +3,7 @@ Copyright (c) 2015, 2017 Jeremy Avigad. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Robert Y. Lewis, Johannes Hölzl, Mario Carneiro, Sébastien Gouëzel -/ +import Mathlib.Data.ENNReal.Lemmas import Mathlib.Topology.Bornology.Constructions import Mathlib.Topology.EMetricSpace.Pi import Mathlib.Topology.MetricSpace.Pseudo.Defs diff --git a/Mathlib/Topology/MetricSpace/ThickenedIndicator.lean b/Mathlib/Topology/MetricSpace/ThickenedIndicator.lean index 1df15588c66b4..b3e9cec910f90 100644 --- a/Mathlib/Topology/MetricSpace/ThickenedIndicator.lean +++ b/Mathlib/Topology/MetricSpace/ThickenedIndicator.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Kalle Kytölä. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kalle Kytölä -/ -import Mathlib.Data.ENNReal.Basic +import Mathlib.Data.ENNReal.Lemmas import Mathlib.Topology.ContinuousMap.Bounded import Mathlib.Topology.MetricSpace.Thickening From 40646ea84cdb8a18541b5ca86e10e2f9c23214d1 Mon Sep 17 00:00:00 2001 From: Hannah Scholz Date: Tue, 22 Oct 2024 16:44:14 +0000 Subject: [PATCH 287/505] feat: lemmas about order on ENat (#17979) --- Mathlib/Data/ENat/Basic.lean | 9 +++++++++ 1 file changed, 9 insertions(+) diff --git a/Mathlib/Data/ENat/Basic.lean b/Mathlib/Data/ENat/Basic.lean index 02639995e5e57..83c5cd47708c2 100644 --- a/Mathlib/Data/ENat/Basic.lean +++ b/Mathlib/Data/ENat/Basic.lean @@ -284,4 +284,13 @@ lemma one_le_iff_ne_zero_withTop {n : WithTop ℕ∞} : ⟨fun h ↦ (zero_lt_one.trans_le h).ne', fun h ↦ add_one_nat_le_withTop_of_lt (pos_iff_ne_zero.mpr h)⟩ +lemma add_one_pos : 0 < n + 1 := + succ_def n ▸ Order.bot_lt_succ n + +lemma add_lt_add_iff_right {k : ℕ∞} (h : k ≠ ⊤) : n + k < m + k ↔ n < m := + WithTop.add_lt_add_iff_right h + +lemma add_lt_add_iff_left {k : ℕ∞} (h : k ≠ ⊤) : k + n < k + m ↔ n < m := + WithTop.add_lt_add_iff_left h + end ENat From ddeafa9973fd347f0eef5111cd16530c7def9331 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 22 Oct 2024 16:44:15 +0000 Subject: [PATCH 288/505] feat(Pointwise): growth of finsets (#18003) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR proves the monotonicity of `n ↦ s ^ n` under various conditions. Also rename `empty_nsmul` to `nsmul_empty`. From GrowthInGroups --- .../Algebra/Group/Pointwise/Finset/Basic.lean | 45 ++++++++++++++++++- .../Algebra/Group/Pointwise/Set/Basic.lean | 19 +++++++- 2 files changed, 60 insertions(+), 4 deletions(-) diff --git a/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean b/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean index 6d68e3fea5280..6e0cef1234889 100644 --- a/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean +++ b/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean @@ -93,6 +93,8 @@ lemma coe_eq_one : (s : Set α) = 1 ↔ s = 1 := coe_eq_singleton theorem one_subset : (1 : Finset α) ⊆ s ↔ (1 : α) ∈ s := singleton_subset_iff +-- TODO: This would be a good simp lemma scoped to `Pointwise`, but it seems `@[simp]` can't be +-- scoped @[to_additive] theorem singleton_one : ({1} : Finset α) = 1 := rfl @@ -967,10 +969,22 @@ theorem mem_pow {a : α} {n : ℕ} : set_option tactic.skipAssignedInstances false in simp [← mem_coe, coe_pow, Set.mem_pow] -@[to_additive (attr := simp)] +@[to_additive (attr := simp) nsmul_empty] theorem empty_pow (hn : n ≠ 0) : (∅ : Finset α) ^ n = ∅ := by rw [← tsub_add_cancel_of_le (Nat.succ_le_of_lt <| Nat.pos_of_ne_zero hn), pow_succ', empty_mul] +@[deprecated (since := "2024-10-21")] alias empty_nsmul := nsmul_empty + +@[to_additive (attr := simp) nsmul_singleton] +lemma singleton_pow (a : α) : ∀ n, ({a} : Finset α) ^ n = {a ^ n} + | 0 => by simp [singleton_one] + | n + 1 => by simp [pow_succ, singleton_pow _ n] + +@[to_additive] +lemma card_pow_le : ∀ {n}, (s ^ n).card ≤ s.card ^ n + | 0 => by simp + | n + 1 => by rw [pow_succ, pow_succ]; refine card_mul_le.trans (by gcongr; exact card_pow_le) + @[to_additive] theorem mul_univ_of_one_mem [Fintype α] (hs : (1 : α) ∈ s) : s * univ = univ := eq_univ_iff_forall.2 fun _ => mem_mul.2 ⟨_, hs, _, mem_univ _, one_mul _⟩ @@ -1015,7 +1029,7 @@ open Pointwise section DivisionMonoid -variable [DivisionMonoid α] {s t : Finset α} +variable [DivisionMonoid α] {s t : Finset α} {n : ℤ} @[to_additive (attr := simp)] theorem coe_zpow (s : Finset α) : ∀ n : ℤ, ↑(s ^ n) = (s : Set α) ^ n @@ -1060,6 +1074,12 @@ lemma univ_div_univ [Fintype α] : (univ / univ : Finset α) = univ := by simp [ @[to_additive] lemma inv_subset_div_right (hs : 1 ∈ s) : t⁻¹ ⊆ s / t := by rw [div_eq_mul_inv]; exact subset_mul_right _ hs +@[to_additive (attr := simp) zsmul_empty] +lemma empty_zpow (hn : n ≠ 0) : (∅ : Finset α) ^ n = ∅ := by cases n <;> aesop + +@[to_additive (attr := simp) zsmul_singleton] +lemma singleton_zpow (a : α) (n : ℤ) : ({a} : Finset α) ^ n = {a ^ n} := by cases n <;> simp + end DivisionMonoid /-- `Finset α` is a commutative division monoid under pointwise operations if `α` is. -/ @@ -1455,6 +1475,27 @@ theorem card_le_card_mul_self' {s : Finset α} : s.card ≤ (s * s).card := by end +section CancelMonoid +variable [DecidableEq α] [CancelMonoid α] {s : Finset α} {m n : ℕ} + +/-- See `Finset.card_pow_mono` for a version that works for the empty set. -/ +@[to_additive "See `Finset.card_nsmul_mono` for a version that works for the empty set."] +protected lemma Nonempty.card_pow_mono (hs : s.Nonempty) : Monotone fun n : ℕ ↦ (s ^ n).card := + monotone_nat_of_le_succ fun n ↦ by rw [pow_succ]; exact card_le_card_mul_right _ hs + +/-- See `Finset.Nonempty.card_pow_mono` for a version that works for zero powers. -/ +@[to_additive "See `Finset.Nonempty.card_nsmul_mono` for a version that works for zero scalars."] +lemma card_pow_mono (hm : m ≠ 0) (hmn : m ≤ n) : (s ^ m).card ≤ (s ^ n).card := by + obtain rfl | hs := s.eq_empty_or_nonempty + · simp [hm] + · exact hs.card_pow_mono hmn + +@[to_additive] +lemma card_le_card_pow (hn : n ≠ 0) : s.card ≤ (s ^ n).card := by + simpa using card_pow_mono (s := s) one_ne_zero (Nat.one_le_iff_ne_zero.2 hn) + +end CancelMonoid + section Group variable [Group α] [DecidableEq α] {s t : Finset α} diff --git a/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean b/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean index bb8aba1b06f3a..946330326b6c8 100644 --- a/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean +++ b/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean @@ -89,6 +89,8 @@ scoped[Pointwise] attribute [instance] Set.one Set.zero open Pointwise +-- TODO: This would be a good simp lemma scoped to `Pointwise`, but it seems `@[simp]` can't be +-- scoped @[to_additive] theorem singleton_one : ({1} : Set α) = 1 := rfl @@ -1098,11 +1100,18 @@ theorem pow_subset_pow_of_one_mem (hs : (1 : α) ∈ s) (hn : m ≤ n) : s ^ m rw [pow_succ'] exact ih.trans (subset_mul_right _ hs) -@[to_additive (attr := simp)] +@[to_additive (attr := simp) nsmul_empty] theorem empty_pow {n : ℕ} (hn : n ≠ 0) : (∅ : Set α) ^ n = ∅ := by match n with | n + 1 => rw [pow_succ', empty_mul] +@[deprecated (since := "2024-10-21")] alias empty_nsmul := nsmul_empty + +@[to_additive (attr := simp) nsmul_singleton] +lemma singleton_pow (a : α) : ∀ n, ({a} : Set α) ^ n = {a ^ n} + | 0 => by simp [singleton_one] + | n + 1 => by simp [pow_succ, singleton_pow _ n] + @[to_additive] theorem mul_univ_of_one_mem (hs : (1 : α) ∈ s) : s * univ = univ := eq_univ_iff_forall.2 fun _ => mem_mul.2 ⟨_, hs, _, mem_univ _, one_mul _⟩ @@ -1138,7 +1147,7 @@ open Pointwise section DivisionMonoid -variable [DivisionMonoid α] {s t : Set α} +variable [DivisionMonoid α] {s t : Set α} {n : ℤ} @[to_additive] protected theorem mul_eq_one_iff : s * t = 1 ↔ ∃ a b, s = {a} ∧ t = {b} ∧ a * b = 1 := by @@ -1191,6 +1200,12 @@ lemma univ_div_univ : (univ / univ : Set α) = univ := by simp [div_eq_mul_inv] @[to_additive] lemma inv_subset_div_right (hs : 1 ∈ s) : t⁻¹ ⊆ s / t := by rw [div_eq_mul_inv]; exact subset_mul_right _ hs +@[to_additive (attr := simp) zsmul_empty] +lemma empty_zpow (hn : n ≠ 0) : (∅ : Set α) ^ n = ∅ := by cases n <;> aesop + +@[to_additive (attr := simp) zsmul_singleton] +lemma singleton_zpow (a : α) (n : ℤ) : ({a} : Set α) ^ n = {a ^ n} := by cases n <;> simp + end DivisionMonoid /-- `Set α` is a commutative division monoid under pointwise operations if `α` is. -/ From 3fc88df14c066088e5e4414114cd3d69a4de8bb2 Mon Sep 17 00:00:00 2001 From: FR Date: Tue, 22 Oct 2024 16:44:16 +0000 Subject: [PATCH 289/505] chore: deprecate `Nat.binaryRec_eq`, rename `Nat.binaryRec_eq'` => `Nat.binaryRec_eq` (#18039) --- Mathlib/Data/Nat/BinaryRec.lean | 13 +++---------- Mathlib/Data/Nat/BitIndices.lean | 4 ++-- Mathlib/Data/Nat/Bits.lean | 2 +- Mathlib/Data/Nat/Bitwise.lean | 2 +- Mathlib/Data/Nat/EvenOddRec.lean | 4 ++-- Mathlib/Data/Num/Lemmas.lean | 11 +++++------ 6 files changed, 14 insertions(+), 22 deletions(-) diff --git a/Mathlib/Data/Nat/BinaryRec.lean b/Mathlib/Data/Nat/BinaryRec.lean index 0edc733fa68b0..a3d3e6f8e1fc1 100644 --- a/Mathlib/Data/Nat/BinaryRec.lean +++ b/Mathlib/Data/Nat/BinaryRec.lean @@ -3,6 +3,7 @@ Copyright (c) 2017 Microsoft Corporation. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Praneeth Kolichala, Yuyang Zhao -/ +import Batteries.Tactic.Alias /-! # Binary recursion on `Nat` @@ -127,12 +128,7 @@ theorem binaryRec_one {motive : Nat → Sort u} (z : motive 0) (f : ∀ b n, mot simp only [add_one_ne_zero, ↓reduceDIte, Nat.reduceShiftRight, binaryRec_zero] rfl -/-- -The same as `binaryRec_eq`, -but that one unfortunately requires `f` to be the identity when appending `false` to `0`. -Here, we allow you to explicitly say that that case is not happening, -i.e. supplying `n = 0 → b = true`. -/ -theorem binaryRec_eq' {motive : Nat → Sort u} {z : motive 0} +theorem binaryRec_eq {motive : Nat → Sort u} {z : motive 0} {f : ∀ b n, motive n → motive (bit b n)} (b n) (h : f false 0 z = z ∨ (n = 0 → b = true)) : binaryRec z f (bit b n) = f b n (binaryRec z f n) := by @@ -149,9 +145,6 @@ theorem binaryRec_eq' {motive : Nat → Sort u} {z : motive 0} rw [testBit_bit_zero, bit_shiftRight_one] intros; rfl -theorem binaryRec_eq {motive : Nat → Sort u} {z : motive 0} {f : ∀ b n, motive n → motive (bit b n)} - (h : f false 0 z = z) (b n) : - binaryRec z f (bit b n) = f b n (binaryRec z f n) := - binaryRec_eq' b n (.inl h) +@[deprecated (since := "2024-10-21")] alias binaryRec_eq' := binaryRec_eq end Nat diff --git a/Mathlib/Data/Nat/BitIndices.lean b/Mathlib/Data/Nat/BitIndices.lean index bd77462f10032..9d9c7c094edad 100644 --- a/Mathlib/Data/Nat/BitIndices.lean +++ b/Mathlib/Data/Nat/BitIndices.lean @@ -41,11 +41,11 @@ def bitIndices (n : ℕ) : List ℕ := theorem bitIndices_bit_true (n : ℕ) : bitIndices (bit true n) = 0 :: ((bitIndices n).map (· + 1)) := - binaryRec_eq rfl _ _ + binaryRec_eq _ _ (.inl rfl) theorem bitIndices_bit_false (n : ℕ) : bitIndices (bit false n) = (bitIndices n).map (· + 1) := - binaryRec_eq rfl _ _ + binaryRec_eq _ _ (.inl rfl) @[simp] theorem bitIndices_two_mul_add_one (n : ℕ) : bitIndices (2 * n + 1) = 0 :: (bitIndices n).map (· + 1) := by diff --git a/Mathlib/Data/Nat/Bits.lean b/Mathlib/Data/Nat/Bits.lean index bdcc8d87a6a7c..7079602da88d9 100644 --- a/Mathlib/Data/Nat/Bits.lean +++ b/Mathlib/Data/Nat/Bits.lean @@ -264,7 +264,7 @@ theorem zero_bits : bits 0 = [] := by simp [Nat.bits] @[simp] theorem bits_append_bit (n : ℕ) (b : Bool) (hn : n = 0 → b = true) : (bit b n).bits = b :: n.bits := by - rw [Nat.bits, Nat.bits, binaryRec_eq'] + rw [Nat.bits, Nat.bits, binaryRec_eq] simpa @[simp] diff --git a/Mathlib/Data/Nat/Bitwise.lean b/Mathlib/Data/Nat/Bitwise.lean index 3c394416fac8e..70965de4413df 100644 --- a/Mathlib/Data/Nat/Bitwise.lean +++ b/Mathlib/Data/Nat/Bitwise.lean @@ -71,7 +71,7 @@ theorem binaryRec_of_ne_zero {C : Nat → Sort*} (z : C 0) (f : ∀ b n, C n → binaryRec z f n = bit_decomp n ▸ f (bodd n) (div2 n) (binaryRec z f (div2 n)) := by cases n using bitCasesOn with | h b n => - rw [binaryRec_eq' _ _ (by right; simpa [bit_eq_zero_iff] using h)] + rw [binaryRec_eq _ _ (by right; simpa [bit_eq_zero_iff] using h)] generalize_proofs h; revert h rw [bodd_bit, div2_bit] simp diff --git a/Mathlib/Data/Nat/EvenOddRec.lean b/Mathlib/Data/Nat/EvenOddRec.lean index ce6c89abcd14f..679a7ff57f097 100644 --- a/Mathlib/Data/Nat/EvenOddRec.lean +++ b/Mathlib/Data/Nat/EvenOddRec.lean @@ -30,14 +30,14 @@ theorem evenOddRec_zero {P : ℕ → Sort*} (h0 : P 0) (h_even : ∀ i, P i → theorem evenOddRec_even {P : ℕ → Sort*} (h0 : P 0) (h_even : ∀ i, P i → P (2 * i)) (h_odd : ∀ i, P i → P (2 * i + 1)) (H : h_even 0 h0 = h0) (n : ℕ) : (2 * n).evenOddRec h0 h_even h_odd = h_even n (evenOddRec h0 h_even h_odd n) := by - apply binaryRec_eq' false n + apply binaryRec_eq false n simp [H] @[simp] theorem evenOddRec_odd {P : ℕ → Sort*} (h0 : P 0) (h_even : ∀ i, P i → P (2 * i)) (h_odd : ∀ i, P i → P (2 * i + 1)) (H : h_even 0 h0 = h0) (n : ℕ) : (2 * n + 1).evenOddRec h0 h_even h_odd = h_odd n (evenOddRec h0 h_even h_odd n) := by - apply binaryRec_eq' true n + apply binaryRec_eq true n simp [H] /-- Strong recursion principle on even and odd numbers: if for all `i : ℕ` we can prove `P (2 * i)` diff --git a/Mathlib/Data/Num/Lemmas.lean b/Mathlib/Data/Num/Lemmas.lean index 2808a6ef83358..1ee383be417b7 100644 --- a/Mathlib/Data/Num/Lemmas.lean +++ b/Mathlib/Data/Num/Lemmas.lean @@ -201,7 +201,7 @@ theorem bit1_of_bit1 : ∀ n : Num, (n + n) + 1 = n.bit1 theorem ofNat'_zero : Num.ofNat' 0 = 0 := by simp [Num.ofNat'] theorem ofNat'_bit (b n) : ofNat' (Nat.bit b n) = cond b Num.bit1 Num.bit0 (ofNat' n) := - Nat.binaryRec_eq rfl _ _ + Nat.binaryRec_eq _ _ (.inl rfl) @[simp] theorem ofNat'_one : Num.ofNat' 1 = 1 := by erw [ofNat'_bit true 0, cond, ofNat'_zero]; rfl @@ -661,12 +661,11 @@ theorem natSize_to_nat (n) : natSize n = Nat.size n := by rw [← size_eq_natSiz theorem ofNat'_eq : ∀ n, Num.ofNat' n = n := Nat.binaryRec (by simp) fun b n IH => by rw [ofNat'] at IH ⊢ - rw [Nat.binaryRec_eq, IH] + rw [Nat.binaryRec_eq _ _ (.inl rfl), IH] -- Porting note: `Nat.cast_bit0` & `Nat.cast_bit1` are not `simp` theorems anymore. - · cases b <;> simp only [cond_false, cond_true, Nat.bit, two_mul, Nat.cast_add, Nat.cast_one] - · rw [bit0_of_bit0] - · rw [bit1_of_bit1] - · rfl + cases b <;> simp only [cond_false, cond_true, Nat.bit, two_mul, Nat.cast_add, Nat.cast_one] + · rw [bit0_of_bit0] + · rw [bit1_of_bit1] theorem zneg_toZNum (n : Num) : -n.toZNum = n.toZNumNeg := by cases n <;> rfl From 2deff082b8dd55ca2a5addc304d8100e1f59ca88 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Tue, 22 Oct 2024 17:17:55 +0000 Subject: [PATCH 290/505] chore(SetTheory/Cardinal/Cofinality): tweak universes of `RelIso.cof_eq_lift` (#18042) For `c : Cardinal.{u}`, the simp-normal form of `lift.{max u v} c` is `lift.{v} c`. Also, some drive-by golfing. --- Mathlib/SetTheory/Cardinal/Cofinality.lean | 30 +++++++--------------- 1 file changed, 9 insertions(+), 21 deletions(-) diff --git a/Mathlib/SetTheory/Cardinal/Cofinality.lean b/Mathlib/SetTheory/Cardinal/Cofinality.lean index e7fb33c1cf155..983ec59006e06 100644 --- a/Mathlib/SetTheory/Cardinal/Cofinality.lean +++ b/Mathlib/SetTheory/Cardinal/Cofinality.lean @@ -79,21 +79,17 @@ theorem le_cof {r : α → α → Prop} [IsRefl α r] (c : Cardinal) : end Order theorem RelIso.cof_le_lift {α : Type u} {β : Type v} {r : α → α → Prop} {s} [IsRefl β s] - (f : r ≃r s) : Cardinal.lift.{max u v} (Order.cof r) ≤ - Cardinal.lift.{max u v} (Order.cof s) := by - rw [Order.cof, Order.cof, lift_sInf, lift_sInf, - le_csInf_iff'' ((Order.cof_nonempty s).image _)] + (f : r ≃r s) : Cardinal.lift.{v} (Order.cof r) ≤ Cardinal.lift.{u} (Order.cof s) := by + rw [Order.cof, Order.cof, lift_sInf, lift_sInf, le_csInf_iff'' ((Order.cof_nonempty s).image _)] rintro - ⟨-, ⟨u, H, rfl⟩, rfl⟩ apply csInf_le' - refine - ⟨_, ⟨f.symm '' u, fun a => ?_, rfl⟩, - lift_mk_eq.{u, v, max u v}.2 ⟨(f.symm.toEquiv.image u).symm⟩⟩ + refine ⟨_, ⟨f.symm '' u, fun a => ?_, rfl⟩, lift_mk_eq'.2 ⟨(f.symm.toEquiv.image u).symm⟩⟩ rcases H (f a) with ⟨b, hb, hb'⟩ refine ⟨f.symm b, mem_image_of_mem _ hb, f.map_rel_iff.1 ?_⟩ rwa [RelIso.apply_symm_apply] theorem RelIso.cof_eq_lift {α : Type u} {β : Type v} {r s} [IsRefl α r] [IsRefl β s] (f : r ≃r s) : - Cardinal.lift.{max u v} (Order.cof r) = Cardinal.lift.{max u v} (Order.cof s) := + Cardinal.lift.{v} (Order.cof r) = Cardinal.lift.{u} (Order.cof s) := (RelIso.cof_le_lift f).antisymm (RelIso.cof_le_lift f.symm) theorem RelIso.cof_le {α β : Type u} {r : α → α → Prop} {s} [IsRefl β s] (f : r ≃r s) : @@ -124,19 +120,11 @@ namespace Ordinal `∀ a, ∃ b ∈ S, a ≤ b`. It is defined for all ordinals, but `cof 0 = 0` and `cof (succ o) = 1`, so it is only really interesting on limit ordinals (when it is an infinite cardinal). -/ -def cof (o : Ordinal.{u}) : Cardinal.{u} := - o.liftOn (fun a => StrictOrder.cof a.r) - (by - rintro ⟨α, r, wo₁⟩ ⟨β, s, wo₂⟩ ⟨⟨f, hf⟩⟩ - haveI := wo₁; haveI := wo₂ - dsimp only - apply @RelIso.cof_eq _ _ _ _ ?_ ?_ - · constructor - exact @fun a b => not_iff_not.2 hf - · dsimp only [swap] - exact ⟨fun _ => irrefl _⟩ - · dsimp only [swap] - exact ⟨fun _ => irrefl _⟩) +def cof (o : Ordinal.{u}) : Cardinal.{u} := by + refine o.liftOn (fun a => StrictOrder.cof a.r) ?_ + rintro ⟨α, r, _⟩ ⟨β, s, _⟩ ⟨⟨f, hf⟩⟩ + refine @RelIso.cof_eq _ _ _ _ ?_ ?_ ⟨_, not_iff_not.2 hf⟩ <;> + exact ⟨fun _ => irrefl _⟩ theorem cof_type (r : α → α → Prop) [IsWellOrder α r] : (type r).cof = StrictOrder.cof r := rfl From 3d42f39dbc59acba8644c9531d4bc9c56e380fd2 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Tue, 22 Oct 2024 17:17:56 +0000 Subject: [PATCH 291/505] =?UTF-8?q?refactor(SetTheory/Ordinal/Basic):=20`t?= =?UTF-8?q?ype=5Flt`=20=E2=86=92=20`type=5FtoType`=20(#18049)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The new name is much more discoverable and makes clear that this isn't just any `<` relation. --- .../MeasureTheory/MeasurableSpace/Card.lean | 2 +- Mathlib/SetTheory/Cardinal/Cofinality.lean | 8 ++--- Mathlib/SetTheory/Ordinal/Arithmetic.lean | 11 +++--- Mathlib/SetTheory/Ordinal/Basic.lean | 35 +++++++++---------- 4 files changed, 28 insertions(+), 28 deletions(-) diff --git a/Mathlib/MeasureTheory/MeasurableSpace/Card.lean b/Mathlib/MeasureTheory/MeasurableSpace/Card.lean index 9da03d77e896e..f6d160c7eefc7 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/Card.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/Card.lean @@ -121,7 +121,7 @@ theorem generateMeasurable_eq_rec (s : Set (Set α)) : refine mem_iUnion.2 ⟨Ordinal.enum (α := ω₁) (· < ·) ⟨Ordinal.lsub fun n => Ordinal.typein.{u} (α := ω₁) (· < ·) (I n), ?_⟩, iUnion_mem_generateMeasurableRec fun n => ⟨I n, ?_, hI n⟩⟩ - · rw [Ordinal.type_lt] + · rw [Ordinal.type_toType] refine Ordinal.lsub_lt_ord_lift ?_ fun i => Ordinal.typein_lt_self _ rw [mk_denumerable, lift_aleph0, isRegular_aleph_one.cof_eq] exact aleph0_lt_aleph_one diff --git a/Mathlib/SetTheory/Cardinal/Cofinality.lean b/Mathlib/SetTheory/Cardinal/Cofinality.lean index 983ec59006e06..e0ff450fad422 100644 --- a/Mathlib/SetTheory/Cardinal/Cofinality.lean +++ b/Mathlib/SetTheory/Cardinal/Cofinality.lean @@ -194,7 +194,7 @@ theorem cof_eq_sInf_lsub (o : Ordinal.{u}) : cof o = sInf { a : Cardinal | ∃ (ι : Type u) (f : ι → Ordinal), lsub.{u, u} f = o ∧ #ι = a } := by refine le_antisymm (le_csInf (cof_lsub_def_nonempty o) ?_) (csInf_le' ?_) · rintro a ⟨ι, f, hf, rfl⟩ - rw [← type_lt o] + rw [← type_toType o] refine (cof_type_le fun a => ?_).trans (@mk_le_of_injective _ _ @@ -208,7 +208,7 @@ theorem cof_eq_sInf_lsub (o : Ordinal.{u}) : cof o = simp_rw [← hf, lt_lsub_iff] at this cases' this with i hi refine ⟨enum (α := o.toType) (· < ·) ⟨f i, ?_⟩, ?_, ?_⟩ - · rw [type_lt, ← hf] + · rw [type_toType, ← hf] apply lt_lsub · rw [mem_preimage, typein_enum] exact mem_range_self i @@ -216,8 +216,8 @@ theorem cof_eq_sInf_lsub (o : Ordinal.{u}) : cof o = · rcases cof_eq (α := o.toType) (· < ·) with ⟨S, hS, hS'⟩ let f : S → Ordinal := fun s => typein LT.lt s.val refine ⟨S, f, le_antisymm (lsub_le fun i => typein_lt_self (o := o) i) - (le_of_forall_lt fun a ha => ?_), by rwa [type_lt o] at hS'⟩ - rw [← type_lt o] at ha + (le_of_forall_lt fun a ha => ?_), by rwa [type_toType o] at hS'⟩ + rw [← type_toType o] at ha rcases hS (enum (· < ·) ⟨a, ha⟩) with ⟨b, hb, hb'⟩ rw [← typein_le_typein, typein_enum] at hb' exact hb'.trans_lt (lt_lsub.{u, u} f ⟨b, hb⟩) diff --git a/Mathlib/SetTheory/Ordinal/Arithmetic.lean b/Mathlib/SetTheory/Ordinal/Arithmetic.lean index 074ad4247ca3a..84df898f4ac34 100644 --- a/Mathlib/SetTheory/Ordinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Ordinal/Arithmetic.lean @@ -317,7 +317,7 @@ instance orderTopToTypeSucc (o : Ordinal) : OrderTop (succ o).toType := @OrderTop.mk _ _ (Top.mk _) le_enum_succ theorem enum_succ_eq_top {o : Ordinal} : - enum (α := (succ o).toType) (· < ·) ⟨o, by rw [type_lt]; exact lt_succ o⟩ = ⊤ := + enum (α := (succ o).toType) (· < ·) ⟨o, type_toType _ ▸ lt_succ o⟩ = ⊤ := rfl theorem has_succ_of_type_succ_lt {α} {r : α → α → Prop} [wo : IsWellOrder α r] @@ -328,7 +328,7 @@ theorem has_succ_of_type_succ_lt {α} {r : α → α → Prop} [wo : IsWellOrder · rw [Subtype.mk_lt_mk, lt_succ_iff] theorem toType_noMax_of_succ_lt {o : Ordinal} (ho : ∀ a < o, succ a < o) : NoMaxOrder o.toType := - ⟨has_succ_of_type_succ_lt (by rwa [type_lt])⟩ + ⟨has_succ_of_type_succ_lt (type_toType _ ▸ ho)⟩ @[deprecated toType_noMax_of_succ_lt (since := "2024-08-26")] alias out_no_max_of_succ_lt := toType_noMax_of_succ_lt @@ -1072,7 +1072,7 @@ def familyOfBFamily' {ι : Type u} (r : ι → ι → Prop) [IsWellOrder ι r] { /-- Converts a family indexed by an `Ordinal.{u}` to one indexed by a `Type u` using a well-ordering given by the axiom of choice. -/ def familyOfBFamily (o : Ordinal) (f : ∀ a < o, α) : o.toType → α := - familyOfBFamily' (· < ·) (type_lt o) f + familyOfBFamily' (· < ·) (type_toType o) f @[simp] theorem bfamilyOfFamily'_typein {ι} (r : ι → ι → Prop) [IsWellOrder ι r] (f : ι → α) (i) : @@ -1092,8 +1092,9 @@ theorem familyOfBFamily'_enum {ι : Type u} (r : ι → ι → Prop) [IsWellOrde @[simp, nolint simpNF] -- Porting note (#10959): simp cannot prove this theorem familyOfBFamily_enum (o : Ordinal) (f : ∀ a < o, α) (i hi) : - familyOfBFamily o f (enum (α := o.toType) (· < ·) ⟨i, hi.trans_eq (type_lt _).symm⟩) = f i hi := - familyOfBFamily'_enum _ (type_lt o) f _ _ + familyOfBFamily o f (enum (α := o.toType) (· < ·) ⟨i, hi.trans_eq (type_toType _).symm⟩) + = f i hi := + familyOfBFamily'_enum _ (type_toType o) f _ _ /-- The range of a family indexed by ordinals. -/ def brange (o : Ordinal) (f : ∀ a < o, α) : Set α := diff --git a/Mathlib/SetTheory/Ordinal/Basic.lean b/Mathlib/SetTheory/Ordinal/Basic.lean index bc2a81d9ed4ff..4c33da600d285 100644 --- a/Mathlib/SetTheory/Ordinal/Basic.lean +++ b/Mathlib/SetTheory/Ordinal/Basic.lean @@ -149,12 +149,16 @@ theorem type_def (r) [wo : IsWellOrder α r] : (⟦⟨α, r, wo⟩⟧ : Ordinal) rfl @[simp] +theorem type_toType (o : Ordinal) : type (α := o.toType) (· < ·) = o := + (type_def' _).symm.trans <| Quotient.out_eq o + +@[deprecated type_toType (since := "2024-10-22")] theorem type_lt (o : Ordinal) : type (α := o.toType) (· < ·) = o := (type_def' _).symm.trans <| Quotient.out_eq o -@[deprecated type_lt (since := "2024-08-26")] +@[deprecated type_toType (since := "2024-08-26")] theorem type_out (o : Ordinal) : Ordinal.type o.out.r = o := - type_lt o + type_toType o theorem type_eq {α β} {r : α → α → Prop} {s : β → β → Prop} [IsWellOrder α r] [IsWellOrder β s] : type r = type s ↔ Nonempty (r ≃r s) := @@ -202,7 +206,7 @@ theorem type_unit : type (@EmptyRelation Unit) = 1 := @[simp] theorem toType_empty_iff_eq_zero {o : Ordinal} : IsEmpty o.toType ↔ o = 0 := by - rw [← @type_eq_zero_iff_isEmpty o.toType (· < ·), type_lt] + rw [← @type_eq_zero_iff_isEmpty o.toType (· < ·), type_toType] @[deprecated toType_empty_iff_eq_zero (since := "2024-08-26")] alias out_empty_iff_eq_zero := toType_empty_iff_eq_zero @@ -216,7 +220,7 @@ instance isEmpty_toType_zero : IsEmpty (toType 0) := @[simp] theorem toType_nonempty_iff_ne_zero {o : Ordinal} : Nonempty o.toType ↔ o ≠ 0 := by - rw [← @type_ne_zero_iff_nonempty o.toType (· < ·), type_lt] + rw [← @type_ne_zero_iff_nonempty o.toType (· < ·), type_toType] @[deprecated toType_nonempty_iff_ne_zero (since := "2024-08-26")] alias out_nonempty_iff_ne_zero := toType_nonempty_iff_ne_zero @@ -356,7 +360,7 @@ instance NeZero.one : NeZero (1 : Ordinal) := `α.toType` into `β.toType`. -/ def initialSegToType {α β : Ordinal} (h : α ≤ β) : α.toType ≤i β.toType := by apply Classical.choice (type_le_iff.mp _) - rwa [type_lt, type_lt] + rwa [type_toType, type_toType] @[deprecated initialSegToType (since := "2024-08-26")] noncomputable alias initialSegOut := initialSegToType @@ -365,7 +369,7 @@ noncomputable alias initialSegOut := initialSegToType of `α.toType` into `β.toType`. -/ def principalSegToType {α β : Ordinal} (h : α < β) : α.toType Subtype.ext_val (typein_enum _ _) right_inv _ := enum_typein _ _ - map_rel_iff' := by - rintro ⟨a, _⟩ ⟨b, _⟩ - apply enum_le_enum' + map_rel_iff' := enum_le_enum' _ @[deprecated (since := "2024-08-26")] alias enumIsoOut := enumIsoToType @@ -999,7 +998,7 @@ def toTypeOrderBotOfPos {o : Ordinal} (ho : 0 < o) : OrderBot o.toType where noncomputable alias outOrderBotOfPos := toTypeOrderBotOfPos theorem enum_zero_eq_bot {o : Ordinal} (ho : 0 < o) : - enum (α := o.toType) (· < ·) ⟨0, by rwa [type_lt]⟩ = + enum (α := o.toType) (· < ·) ⟨0, by rwa [type_toType]⟩ = have _ := toTypeOrderBotOfPos ho (⊥ : o.toType) := rfl @@ -1097,7 +1096,7 @@ open Ordinal @[simp] theorem mk_toType (o : Ordinal) : #o.toType = o.card := - (Ordinal.card_type _).symm.trans <| by rw [Ordinal.type_lt] + (Ordinal.card_type _).symm.trans <| by rw [Ordinal.type_toType] @[deprecated mk_toType (since := "2024-08-26")] alias mk_ordinal_out := mk_toType From 877aae73ce134e68b43d24adeab9175e218a1be4 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Tue, 22 Oct 2024 17:17:57 +0000 Subject: [PATCH 292/505] chore(SetTheory/Ordinal/Rank): move `IsWellFounded.rank` to its own file (#18078) --- Mathlib.lean | 1 + Mathlib/Order/Extension/Well.lean | 2 +- Mathlib/SetTheory/Ordinal/Arithmetic.lean | 83 ------------------ Mathlib/SetTheory/Ordinal/Rank.lean | 102 ++++++++++++++++++++++ Mathlib/SetTheory/ZFC/Rank.lean | 2 +- 5 files changed, 105 insertions(+), 85 deletions(-) create mode 100644 Mathlib/SetTheory/Ordinal/Rank.lean diff --git a/Mathlib.lean b/Mathlib.lean index fe64af1dced4a..408df6cd2784b 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4256,6 +4256,7 @@ import Mathlib.SetTheory.Ordinal.NaturalOps import Mathlib.SetTheory.Ordinal.Nimber import Mathlib.SetTheory.Ordinal.Notation import Mathlib.SetTheory.Ordinal.Principal +import Mathlib.SetTheory.Ordinal.Rank import Mathlib.SetTheory.Ordinal.Topology import Mathlib.SetTheory.Surreal.Basic import Mathlib.SetTheory.Surreal.Dyadic diff --git a/Mathlib/Order/Extension/Well.lean b/Mathlib/Order/Extension/Well.lean index f3f4d25022936..1b5175e593d0c 100644 --- a/Mathlib/Order/Extension/Well.lean +++ b/Mathlib/Order/Extension/Well.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies, Junyan Xu -/ import Mathlib.Data.Prod.Lex -import Mathlib.SetTheory.Ordinal.Arithmetic +import Mathlib.SetTheory.Ordinal.Rank /-! # Extend a well-founded order to a well-order diff --git a/Mathlib/SetTheory/Ordinal/Arithmetic.lean b/Mathlib/SetTheory/Ordinal/Arithmetic.lean index 84df898f4ac34..491d756f608e1 100644 --- a/Mathlib/SetTheory/Ordinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Ordinal/Arithmetic.lean @@ -2416,87 +2416,4 @@ theorem noMaxOrder {c} (h : ℵ₀ ≤ c) : NoMaxOrder c.ord.toType := end Cardinal -variable {α : Type u} {a b : α} - -namespace Acc - -variable {r : α → α → Prop} - -/-- The rank of an element `a` accessible under a relation `r` is defined inductively as the -smallest ordinal greater than the ranks of all elements below it (i.e. elements `b` such that -`r b a`). -/ -noncomputable def rank (h : Acc r a) : Ordinal.{u} := - Acc.recOn h fun a _h ih => ⨆ b : { b // r b a }, Order.succ (ih b b.2) - -theorem rank_eq (h : Acc r a) : - h.rank = ⨆ b : { b // r b a }, Order.succ (h.inv b.2).rank := by - change (Acc.intro a fun _ => h.inv).rank = _ - rfl - -/-- if `r a b` then the rank of `a` is less than the rank of `b`. -/ -theorem rank_lt_of_rel (hb : Acc r b) (h : r a b) : (hb.inv h).rank < hb.rank := - (Order.lt_succ _).trans_le <| by - rw [hb.rank_eq] - exact Ordinal.le_iSup _ (⟨a, h⟩ : {a // r a b}) - -end Acc - -namespace IsWellFounded - -variable (r : α → α → Prop) [hwf : IsWellFounded α r] - -/-- The rank of an element `a` under a well-founded relation `r` is defined inductively as the -smallest ordinal greater than the ranks of all elements below it (i.e. elements `b` such that -`r b a`). -/ -noncomputable def rank (a : α) : Ordinal.{u} := - (hwf.apply r a).rank - -theorem rank_eq (a : α) : rank r a = ⨆ b : { b // r b a }, succ (rank r b) := - (hwf.apply r a).rank_eq - -variable {r} in -theorem rank_lt_of_rel (h : r a b) : rank r a < rank r b := - Acc.rank_lt_of_rel _ h - -end IsWellFounded - -theorem WellFoundedLT.rank_strictMono [Preorder α] [WellFoundedLT α] : - StrictMono (IsWellFounded.rank (α := α) (· < ·)) := - fun _ _ => IsWellFounded.rank_lt_of_rel - -theorem WellFoundedGT.rank_strictAnti [Preorder α] [WellFoundedGT α] : - StrictAnti (IsWellFounded.rank (α := α) (· > ·)) := - fun _ _ a => IsWellFounded.rank_lt_of_rel a - -namespace WellFounded - -set_option linter.deprecated false - -variable {r : α → α → Prop} (hwf : WellFounded r) - -/-- The rank of an element `a` under a well-founded relation `r` is defined inductively as the -smallest ordinal greater than the ranks of all elements below it (i.e. elements `b` such that -`r b a`). -/ -@[deprecated IsWellFounded.rank (since := "2024-09-07")] -noncomputable def rank (a : α) : Ordinal.{u} := - (hwf.apply a).rank - -@[deprecated IsWellFounded.rank_eq (since := "2024-09-07")] -theorem rank_eq : hwf.rank a = ⨆ b : { b // r b a }, Order.succ (hwf.rank b) := - (hwf.apply a).rank_eq - -@[deprecated IsWellFounded.rank_lt_of_rel (since := "2024-09-07")] -theorem rank_lt_of_rel (h : r a b) : hwf.rank a < hwf.rank b := - Acc.rank_lt_of_rel _ h - -@[deprecated WellFoundedLT.rank_strictMono (since := "2024-09-07")] -theorem rank_strictMono [Preorder α] [WellFoundedLT α] : - StrictMono (rank <| @wellFounded_lt α _ _) := fun _ _ => rank_lt_of_rel _ - -@[deprecated WellFoundedGT.rank_strictAnti (since := "2024-09-07")] -theorem rank_strictAnti [Preorder α] [WellFoundedGT α] : - StrictAnti (rank <| @wellFounded_gt α _ _) := fun _ _ => rank_lt_of_rel wellFounded_gt - -end WellFounded - set_option linter.style.longFile 2600 diff --git a/Mathlib/SetTheory/Ordinal/Rank.lean b/Mathlib/SetTheory/Ordinal/Rank.lean new file mode 100644 index 0000000000000..20aa63340bd04 --- /dev/null +++ b/Mathlib/SetTheory/Ordinal/Rank.lean @@ -0,0 +1,102 @@ +/- +Copyright (c) 2022 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies +-/ +import Mathlib.SetTheory.Ordinal.Arithmetic + +/-! +# Rank in a well-founded relation + +For `r` a well-founded relation, `IsWellFounded.rank r a` is recursively defined as the least +ordinal greater than the ranks of all elements below `a`. +-/ + +universe u + +variable {α : Type u} {a b : α} + +/-! ### Rank of an accessible value -/ + +namespace Acc + +variable {r : α → α → Prop} + +/-- The rank of an element `a` accessible under a relation `r` is defined inductively as the +smallest ordinal greater than the ranks of all elements below it (i.e. elements `b` such that +`r b a`). -/ +noncomputable def rank (h : Acc r a) : Ordinal.{u} := + Acc.recOn h fun a _h ih => ⨆ b : { b // r b a }, Order.succ (ih b b.2) + +theorem rank_eq (h : Acc r a) : + h.rank = ⨆ b : { b // r b a }, Order.succ (h.inv b.2).rank := by + change (Acc.intro a fun _ => h.inv).rank = _ + rfl + +/-- if `r a b` then the rank of `a` is less than the rank of `b`. -/ +theorem rank_lt_of_rel (hb : Acc r b) (h : r a b) : (hb.inv h).rank < hb.rank := + (Order.lt_succ _).trans_le <| by + rw [hb.rank_eq] + exact Ordinal.le_iSup _ (⟨a, h⟩ : {a // r a b}) + +end Acc + +/-! ### Rank in a well-founded relation -/ + +namespace IsWellFounded + +variable (r : α → α → Prop) [hwf : IsWellFounded α r] + +/-- The rank of an element `a` under a well-founded relation `r` is defined inductively as the +smallest ordinal greater than the ranks of all elements below it (i.e. elements `b` such that +`r b a`). -/ +noncomputable def rank (a : α) : Ordinal.{u} := + (hwf.apply r a).rank + +theorem rank_eq (a : α) : rank r a = ⨆ b : { b // r b a }, Order.succ (rank r b) := + (hwf.apply r a).rank_eq + +variable {r} in +theorem rank_lt_of_rel (h : r a b) : rank r a < rank r b := + Acc.rank_lt_of_rel _ h + +end IsWellFounded + +theorem WellFoundedLT.rank_strictMono [Preorder α] [WellFoundedLT α] : + StrictMono (IsWellFounded.rank (α := α) (· < ·)) := + fun _ _ => IsWellFounded.rank_lt_of_rel + +theorem WellFoundedGT.rank_strictAnti [Preorder α] [WellFoundedGT α] : + StrictAnti (IsWellFounded.rank (α := α) (· > ·)) := + fun _ _ a => IsWellFounded.rank_lt_of_rel a + +namespace WellFounded + +set_option linter.deprecated false + +variable {r : α → α → Prop} (hwf : WellFounded r) + +/-- The rank of an element `a` under a well-founded relation `r` is defined inductively as the +smallest ordinal greater than the ranks of all elements below it (i.e. elements `b` such that +`r b a`). -/ +@[deprecated IsWellFounded.rank (since := "2024-09-07")] +noncomputable def rank (a : α) : Ordinal.{u} := + (hwf.apply a).rank + +@[deprecated IsWellFounded.rank_eq (since := "2024-09-07")] +theorem rank_eq : hwf.rank a = ⨆ b : { b // r b a }, Order.succ (hwf.rank b) := + (hwf.apply a).rank_eq + +@[deprecated IsWellFounded.rank_lt_of_rel (since := "2024-09-07")] +theorem rank_lt_of_rel (h : r a b) : hwf.rank a < hwf.rank b := + Acc.rank_lt_of_rel _ h + +@[deprecated WellFoundedLT.rank_strictMono (since := "2024-09-07")] +theorem rank_strictMono [Preorder α] [WellFoundedLT α] : + StrictMono (rank <| @wellFounded_lt α _ _) := fun _ _ => rank_lt_of_rel _ + +@[deprecated WellFoundedGT.rank_strictAnti (since := "2024-09-07")] +theorem rank_strictAnti [Preorder α] [WellFoundedGT α] : + StrictAnti (rank <| @wellFounded_gt α _ _) := fun _ _ => rank_lt_of_rel wellFounded_gt + +end WellFounded diff --git a/Mathlib/SetTheory/ZFC/Rank.lean b/Mathlib/SetTheory/ZFC/Rank.lean index fe8e38cbca3b7..cdcacbd303f4e 100644 --- a/Mathlib/SetTheory/ZFC/Rank.lean +++ b/Mathlib/SetTheory/ZFC/Rank.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 Dexin Zhang. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Dexin Zhang -/ -import Mathlib.SetTheory.Ordinal.Arithmetic +import Mathlib.SetTheory.Ordinal.Rank import Mathlib.SetTheory.ZFC.Basic /-! From 94cbdf2a85f7eca27e4515800907bf7ea6bad136 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Tue, 22 Oct 2024 17:17:59 +0000 Subject: [PATCH 293/505] chore(GroupTheory): delete "`simp` can prove this" porting notes (#18083) This PR checks all porting notes of the form "`simp` can prove this" and if this is indeed the case (and the proof appears to make sense), removes the associated porting note. --- Mathlib/GroupTheory/Coset/Basic.lean | 2 +- Mathlib/GroupTheory/FreeAbelianGroup.lean | 1 - Mathlib/GroupTheory/FreeGroup/Basic.lean | 2 -- Mathlib/GroupTheory/MonoidLocalization/Basic.lean | 3 --- Mathlib/GroupTheory/OrderOfElement.lean | 6 +++--- Mathlib/GroupTheory/Perm/Support.lean | 4 ---- 6 files changed, 4 insertions(+), 14 deletions(-) diff --git a/Mathlib/GroupTheory/Coset/Basic.lean b/Mathlib/GroupTheory/Coset/Basic.lean index c5df256757ef5..1b6d848534e71 100644 --- a/Mathlib/GroupTheory/Coset/Basic.lean +++ b/Mathlib/GroupTheory/Coset/Basic.lean @@ -424,7 +424,7 @@ protected theorem eq {a b : α} : (a : α ⧸ s) = b ↔ a⁻¹ * b ∈ s := @[to_additive (attr := deprecated (since := "2024-08-04"))] alias eq' := QuotientGroup.eq -@[to_additive] -- Porting note (#10618): `simp` can prove this. +@[to_additive] theorem out_eq' (a : α ⧸ s) : mk a.out' = a := Quotient.out_eq' a diff --git a/Mathlib/GroupTheory/FreeAbelianGroup.lean b/Mathlib/GroupTheory/FreeAbelianGroup.lean index 2bbbd09ab5770..a6dae650248a8 100644 --- a/Mathlib/GroupTheory/FreeAbelianGroup.lean +++ b/Mathlib/GroupTheory/FreeAbelianGroup.lean @@ -211,7 +211,6 @@ protected theorem map_sub (f : α → β) (x y : FreeAbelianGroup α) : theorem map_of (f : α → β) (y : α) : f <$> of y = of (f y) := rfl --- @[simp] -- Porting note (#10618): simp can prove this theorem pure_bind (f : α → FreeAbelianGroup β) (x) : pure x >>= f = f x := lift.of _ _ diff --git a/Mathlib/GroupTheory/FreeGroup/Basic.lean b/Mathlib/GroupTheory/FreeGroup/Basic.lean index fea26c87878b2..a6694cbd18425 100644 --- a/Mathlib/GroupTheory/FreeGroup/Basic.lean +++ b/Mathlib/GroupTheory/FreeGroup/Basic.lean @@ -838,7 +838,6 @@ protected theorem induction_on {C : FreeGroup α → Prop} (z : FreeGroup α) (C Quot.inductionOn z fun L => List.recOn L C1 fun ⟨x, b⟩ _tl ih => Bool.recOn b (Cm _ _ (Ci _ <| Cp x) ih) (Cm _ _ (Cp x) ih) --- porting note (#10618): simp can prove this: by simp only [@map_pure] @[to_additive] theorem map_pure (f : α → β) (x : α) : f <$> (pure x : FreeGroup α) = pure (f x) := map.of @@ -855,7 +854,6 @@ theorem map_mul (f : α → β) (x y : FreeGroup α) : f <$> (x * y) = f <$> x * theorem map_inv (f : α → β) (x : FreeGroup α) : f <$> x⁻¹ = (f <$> x)⁻¹ := (map f).map_inv x --- porting note (#10618): simp can prove this: by simp only [@pure_bind] @[to_additive] theorem pure_bind (f : α → FreeGroup β) (x) : pure x >>= f = f x := lift.of diff --git a/Mathlib/GroupTheory/MonoidLocalization/Basic.lean b/Mathlib/GroupTheory/MonoidLocalization/Basic.lean index 5c49b43e602bb..6ad962c53ad34 100644 --- a/Mathlib/GroupTheory/MonoidLocalization/Basic.lean +++ b/Mathlib/GroupTheory/MonoidLocalization/Basic.lean @@ -1067,7 +1067,6 @@ theorem mulEquivOfLocalizations_right_inv (k : LocalizationMap S P) : f.ofMulEquivOfLocalizations (f.mulEquivOfLocalizations k) = k := toMap_injective <| f.lift_comp k.map_units --- @[simp] -- Porting note (#10618): simp can prove this @[to_additive addEquivOfLocalizations_right_inv_apply] theorem mulEquivOfLocalizations_right_inv_apply {k : LocalizationMap S P} {x} : (f.ofMulEquivOfLocalizations (f.mulEquivOfLocalizations k)).toMap x = k.toMap x := by simp @@ -1077,7 +1076,6 @@ theorem mulEquivOfLocalizations_left_inv (k : N ≃* P) : f.mulEquivOfLocalizations (f.ofMulEquivOfLocalizations k) = k := DFunLike.ext _ _ fun x ↦ DFunLike.ext_iff.1 (f.lift_of_comp k.toMonoidHom) x --- @[simp] -- Porting note (#10618): simp can prove this @[to_additive] theorem mulEquivOfLocalizations_left_inv_apply {k : N ≃* P} (x) : f.mulEquivOfLocalizations (f.ofMulEquivOfLocalizations k) x = k x := by simp @@ -1269,7 +1267,6 @@ theorem mulEquivOfQuotient_mk' (x y) : mulEquivOfQuotient f ((monoidOf S).mk' x theorem mulEquivOfQuotient_mk (x y) : mulEquivOfQuotient f (mk x y) = f.mk' x y := by rw [mk_eq_monoidOf_mk'_apply]; exact mulEquivOfQuotient_mk' _ _ --- @[simp] -- Porting note (#10618): simp can prove this @[to_additive] theorem mulEquivOfQuotient_monoidOf (x) : mulEquivOfQuotient f ((monoidOf S).toMap x) = f.toMap x := by simp diff --git a/Mathlib/GroupTheory/OrderOfElement.lean b/Mathlib/GroupTheory/OrderOfElement.lean index 1d4e34b53c52d..12dbd230d8ffa 100644 --- a/Mathlib/GroupTheory/OrderOfElement.lean +++ b/Mathlib/GroupTheory/OrderOfElement.lean @@ -574,7 +574,7 @@ theorem orderOf_inv (x : G) : orderOf x⁻¹ = orderOf x := by simp [orderOf_eq_ namespace Subgroup variable {H : Subgroup G} -@[to_additive (attr := norm_cast)] -- Porting note (#10618): simp can prove this (so removed simp) +@[to_additive (attr := norm_cast)] lemma orderOf_coe (a : H) : orderOf (a : G) = orderOf a := orderOf_injective H.subtype Subtype.coe_injective _ @@ -942,11 +942,11 @@ noncomputable def powCoprime {G : Type*} [Group G] (h : (Nat.card G).Coprime n) rwa [zpow_add, zpow_mul, zpow_mul', zpow_natCast, zpow_natCast, zpow_natCast, h.gcd_eq_one, pow_one, pow_card_eq_one', one_zpow, one_mul, eq_comm] at key -@[to_additive] -- Porting note (#10618): simp can prove this (so removed simp) +@[to_additive] theorem powCoprime_one {G : Type*} [Group G] (h : (Nat.card G).Coprime n) : powCoprime h 1 = 1 := one_pow n -@[to_additive] -- Porting note (#10618): simp can prove this (so removed simp) +@[to_additive] theorem powCoprime_inv {G : Type*} [Group G] (h : (Nat.card G).Coprime n) {g : G} : powCoprime h g⁻¹ = (powCoprime h g)⁻¹ := inv_pow g n diff --git a/Mathlib/GroupTheory/Perm/Support.lean b/Mathlib/GroupTheory/Perm/Support.lean index ea8f7d275ba90..06dbc62cdeb17 100644 --- a/Mathlib/GroupTheory/Perm/Support.lean +++ b/Mathlib/GroupTheory/Perm/Support.lean @@ -328,7 +328,6 @@ theorem support_pow_le (σ : Perm α) (n : ℕ) : (σ ^ n).support ≤ σ.suppor theorem support_inv (σ : Perm α) : support σ⁻¹ = σ.support := by simp_rw [Finset.ext_iff, mem_support, not_iff_not, inv_eq_iff_eq.trans eq_comm, imp_true_iff] --- @[simp] -- Porting note (#10618): simp can prove this theorem apply_mem_support {x : α} : f x ∈ f.support ↔ x ∈ f.support := by rw [mem_support, mem_support, Ne, Ne, apply_eq_iff_eq] @@ -363,11 +362,9 @@ lemma ofSubtype_eq_iff {g c : Equiv.Perm α} {s : Finset α} · rw [ofSubtype_apply_of_not_mem (p := (· ∈ s)) _ ha, eq_comm, ← not_mem_support] exact Finset.not_mem_mono hc ha --- @[simp] -- Porting note (#10618): simp can prove this theorem pow_apply_mem_support {n : ℕ} {x : α} : (f ^ n) x ∈ f.support ↔ x ∈ f.support := by simp only [mem_support, ne_eq, apply_pow_apply_eq_iff] --- @[simp] -- Porting note (#10618): simp can prove this theorem zpow_apply_mem_support {n : ℤ} {x : α} : (f ^ n) x ∈ f.support ↔ x ∈ f.support := by simp only [mem_support, ne_eq, apply_zpow_apply_eq_iff] @@ -535,7 +532,6 @@ end ExtendDomain section Card --- @[simp] -- Porting note (#10618): simp can prove thisrove this theorem card_support_eq_zero {f : Perm α} : f.support.card = 0 ↔ f = 1 := by rw [Finset.card_eq_zero, support_eq_empty_iff] From bbab98c8bbe50bba6f434b4e7d8bd9f13e399a3a Mon Sep 17 00:00:00 2001 From: YnirPaz Date: Tue, 22 Oct 2024 17:45:47 +0000 Subject: [PATCH 294/505] feat(SetTheory/Ordinal/Arithmetic): `IsNormal.iSup` (#16866) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Prove `IsNormal.iSup`, the `iSup` version of IsNormal.sup. Co-authored-by: Violeta Hernández --- Mathlib/SetTheory/Cardinal/UnivLE.lean | 8 +++ Mathlib/SetTheory/Ordinal/Arithmetic.lean | 78 +++++++++++++++++------ 2 files changed, 67 insertions(+), 19 deletions(-) diff --git a/Mathlib/SetTheory/Cardinal/UnivLE.lean b/Mathlib/SetTheory/Cardinal/UnivLE.lean index b5241ecd1cddb..dbc592b08c9b4 100644 --- a/Mathlib/SetTheory/Cardinal/UnivLE.lean +++ b/Mathlib/SetTheory/Cardinal/UnivLE.lean @@ -26,6 +26,14 @@ theorem univLE_iff_cardinal_le : UnivLE.{u, v} ↔ univ.{u, v+1} ≤ univ.{v, u+ rw [univ_umax.{v,u}, ← lift_le.{u+1}, lift_univ, lift_lift] exact h.le +theorem univLE_iff_exists_embedding : UnivLE.{u, v} ↔ Nonempty (Ordinal.{u} ↪ Ordinal.{v}) := by + rw [univLE_iff_cardinal_le] + exact lift_mk_le' + +theorem Ordinal.univLE_of_injective {f : Ordinal.{u} → Ordinal.{v}} (h : f.Injective) : + UnivLE.{u, v} := + univLE_iff_exists_embedding.2 ⟨f, h⟩ + /-- Together with transitivity, this shows UnivLE "IsTotalPreorder". -/ theorem univLE_total : UnivLE.{u, v} ∨ UnivLE.{v, u} := by simp_rw [univLE_iff_cardinal_le]; apply le_total diff --git a/Mathlib/SetTheory/Ordinal/Arithmetic.lean b/Mathlib/SetTheory/Ordinal/Arithmetic.lean index 491d756f608e1..058aaef06fd21 100644 --- a/Mathlib/SetTheory/Ordinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Ordinal/Arithmetic.lean @@ -6,7 +6,7 @@ Authors: Mario Carneiro, Floris van Doorn, Violeta Hernández Palacios import Mathlib.SetTheory.Ordinal.Basic import Mathlib.Data.Nat.SuccPred import Mathlib.Algebra.GroupWithZero.Divisibility -import Mathlib.Logic.UnivLE +import Mathlib.SetTheory.Cardinal.UnivLE /-! # Ordinal arithmetic @@ -263,6 +263,16 @@ theorem zero_or_succ_or_limit (o : Ordinal) : o = 0 ∨ (∃ a, o = succ a) ∨ if h : ∃ a, o = succ a then Or.inr (Or.inl h) else Or.inr <| Or.inr ⟨o0, fun _a => (succ_lt_of_not_succ h).2⟩ +-- TODO: this is an iff with `IsSuccPrelimit` +theorem IsLimit.sSup_Iio {o : Ordinal} (h : IsLimit o) : sSup (Iio o) = o := by + apply (csSup_le' (fun a ha ↦ le_of_lt ha)).antisymm + apply le_of_forall_lt + intro a ha + exact (lt_succ a).trans_le (le_csSup bddAbove_Iio (h.succ_lt ha)) + +theorem IsLimit.iSup_Iio {o : Ordinal} (h : IsLimit o) : ⨆ a : Iio o, a.1 = o := by + rw [← sSup_eq_iSup', h.sSup_Iio] + /-- Main induction principle of ordinals: if one can prove a property by induction at successor ordinals and at limit ordinals, then it holds for all ordinals. -/ @[elab_as_elim] @@ -1185,6 +1195,16 @@ theorem bddAbove_iff_small {s : Set Ordinal.{u}} : BddAbove s ↔ Small.{u} s := ⟨fun ⟨a, h⟩ => small_subset <| show s ⊆ Iic a from fun _ hx => h hx, fun _ => bddAbove_of_small _⟩ +theorem bddAbove_image {s : Set Ordinal.{u}} (hf : BddAbove s) + (f : Ordinal.{u} → Ordinal.{max u v}) : BddAbove (f '' s) := by + rw [bddAbove_iff_small] at hf ⊢ + exact small_lift _ + +theorem bddAbove_range_comp {ι : Type u} {f : ι → Ordinal.{v}} (hf : BddAbove (range f)) + (g : Ordinal.{v} → Ordinal.{max v w}) : BddAbove (range (g ∘ f)) := by + rw [range_comp] + exact bddAbove_image hf g + /-- `le_ciSup` whenever the input type is small in the output universe. This lemma sometimes fails to infer `f` in simple cases and needs it to be given explicitly. -/ protected theorem le_iSup {ι} (f : ι → Ordinal.{u}) [Small.{u} ι] : ∀ i, f i ≤ iSup f := @@ -1270,24 +1290,6 @@ theorem sup_eq_zero_iff {ι : Type u} {f : ι → Ordinal.{max u v}} : rw [← Ordinal.le_zero, ← h] exact le_sup f i --- TODO: generalize universes, make sSup version. -theorem IsNormal.map_iSup {f : Ordinal.{max u v} → Ordinal.{max u w}} (H : IsNormal f) {ι : Type u} - (g : ι → Ordinal.{max u v}) [Nonempty ι] : f (⨆ i, g i) = ⨆ i, f (g i) := by - apply eq_of_forall_ge_iff - intro a - rw [H.le_set' Set.univ Set.univ_nonempty g] - · rw [Ordinal.iSup_le_iff] - simp - · intro o - rw [Ordinal.iSup_le_iff] - simp - -set_option linter.deprecated false in -@[deprecated IsNormal.map_iSup (since := "2024-08-27")] -theorem IsNormal.sup {f : Ordinal.{max u v} → Ordinal.{max u w}} (H : IsNormal f) {ι : Type u} - (g : ι → Ordinal.{max u v}) [Nonempty ι] : f (sup.{_, v} g) = sup.{_, w} (f ∘ g) := - H.map_iSup g - set_option linter.deprecated false in @[deprecated ciSup_of_empty (since := "2024-08-27")] theorem sup_empty {ι} [IsEmpty ι] (f : ι → Ordinal) : sup f = 0 := @@ -1367,6 +1369,44 @@ theorem le_sup_shrink_equiv {s : Set Ordinal.{u}} (hs : Small.{u} s) (a) (ha : a convert le_sup.{u, u} (fun x => ((@equivShrink s hs).symm x).val) ((@equivShrink s hs) ⟨a, ha⟩) rw [symm_apply_apply] +theorem IsNormal.map_iSup_of_bddAbove {f : Ordinal.{u} → Ordinal.{v}} (H : IsNormal f) + {ι : Type*} (g : ι → Ordinal.{u}) (hg : BddAbove (range g)) + [Nonempty ι] : f (⨆ i, g i) = ⨆ i, f (g i) := eq_of_forall_ge_iff fun a ↦ by + have := bddAbove_iff_small.mp hg + have := univLE_of_injective H.strictMono.injective + have := Small.trans_univLE.{u, v} (range g) + have hfg : BddAbove (range (f ∘ g)) := bddAbove_iff_small.mpr <| by + rw [range_comp] + exact small_image f (range g) + change _ ↔ ⨆ i, (f ∘ g) i ≤ a + rw [ciSup_le_iff hfg, H.le_set' _ Set.univ_nonempty g] <;> simp [ciSup_le_iff hg] + +theorem IsNormal.map_iSup {f : Ordinal.{u} → Ordinal.{v}} (H : IsNormal f) + {ι : Type w} (g : ι → Ordinal.{u}) [Small.{u} ι] [Nonempty ι] : + f (⨆ i, g i) = ⨆ i, f (g i) := + H.map_iSup_of_bddAbove g (bddAbove_of_small _) + +theorem IsNormal.map_sSup_of_bddAbove {f : Ordinal.{u} → Ordinal.{v}} (H : IsNormal f) + {s : Set Ordinal.{u}} (hs : BddAbove s) (hn : s.Nonempty) : f (sSup s) = sSup (f '' s) := by + have := hn.to_subtype + rw [sSup_eq_iSup', sSup_image', H.map_iSup_of_bddAbove] + rwa [Subtype.range_coe_subtype, setOf_mem_eq] + +theorem IsNormal.map_sSup {f : Ordinal.{u} → Ordinal.{v}} (H : IsNormal f) + {s : Set Ordinal.{u}} (hn : s.Nonempty) [Small.{u} s] : f (sSup s) = sSup (f '' s) := + H.map_sSup_of_bddAbove (bddAbove_of_small s) hn + +set_option linter.deprecated false in +@[deprecated IsNormal.map_iSup (since := "2024-08-27")] +theorem IsNormal.sup {f : Ordinal.{max u v} → Ordinal.{max u w}} (H : IsNormal f) {ι : Type u} + (g : ι → Ordinal.{max u v}) [Nonempty ι] : f (sup.{_, v} g) = sup.{_, w} (f ∘ g) := + H.map_iSup g + +theorem IsNormal.apply_of_isLimit {f : Ordinal.{u} → Ordinal.{v}} (H : IsNormal f) {o : Ordinal} + (ho : IsLimit o) : f o = ⨆ a : Iio o, f a := by + have : Nonempty (Iio o) := ⟨0, ho.pos⟩ + rw [← H.map_iSup, ho.iSup_Iio] + set_option linter.deprecated false in @[deprecated (since := "2024-08-27")] theorem sup_eq_sSup {s : Set Ordinal.{u}} (hs : Small.{u} s) : From a8d917066fcf79263b30686ded10249b1ada4eed Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Tue, 22 Oct 2024 17:45:49 +0000 Subject: [PATCH 295/505] chore(SetTheory/Ordinal/Basic): golf `Ordinal.lift` lemmas (#17517) We define `Ordinal.liftInitialSeg` earlier and use it to more easily prove the other theorems on `lift`. --- Mathlib/SetTheory/Ordinal/Arithmetic.lean | 4 +- Mathlib/SetTheory/Ordinal/Basic.lean | 126 +++++++++++----------- Mathlib/SetTheory/ZFC/Rank.lean | 4 +- 3 files changed, 67 insertions(+), 67 deletions(-) diff --git a/Mathlib/SetTheory/Ordinal/Arithmetic.lean b/Mathlib/SetTheory/Ordinal/Arithmetic.lean index 058aaef06fd21..9725bed77ccd6 100644 --- a/Mathlib/SetTheory/Ordinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Ordinal/Arithmetic.lean @@ -181,7 +181,7 @@ theorem pred_le {a b} : pred a ≤ b ↔ a ≤ succ b := @[simp] theorem lift_is_succ {o : Ordinal.{v}} : (∃ a, lift.{u} o = succ a) ↔ ∃ a, o = succ a := ⟨fun ⟨a, h⟩ => - let ⟨b, e⟩ := lift_down <| show a ≤ lift.{u} o from le_of_lt <| h.symm ▸ lt_succ a + let ⟨b, e⟩ := mem_range_lift_of_le <| show a ≤ lift.{u} o from le_of_lt <| h.symm ▸ lt_succ a ⟨b, (lift_inj.{u,v}).1 <| by rw [h, ← e, lift_succ]⟩, fun ⟨a, h⟩ => ⟨lift.{u} a, by simp only [h, lift_succ]⟩⟩ @@ -242,7 +242,7 @@ theorem lift_isLimit (o : Ordinal.{v}) : IsLimit (lift.{u,v} o) ↔ IsLimit o := and_congr (not_congr <| by simpa only [lift_zero] using @lift_inj o 0) ⟨fun H a h => (lift_lt.{u,v}).1 <| by simpa only [lift_succ] using H _ (lift_lt.2 h), fun H a h => by - obtain ⟨a', rfl⟩ := lift_down h.le + obtain ⟨a', rfl⟩ := mem_range_lift_of_le h.le rw [← lift_succ, lift_lt] exact H a' (lift_lt.1 h)⟩ diff --git a/Mathlib/SetTheory/Ordinal/Basic.lean b/Mathlib/SetTheory/Ordinal/Basic.lean index 4c33da600d285..5c5d1915ecefc 100644 --- a/Mathlib/SetTheory/Ordinal/Basic.lean +++ b/Mathlib/SetTheory/Ordinal/Basic.lean @@ -600,13 +600,6 @@ theorem lift_id : ∀ a, lift.{u, u} a = a := theorem lift_uzero (a : Ordinal.{u}) : lift.{0} a = a := lift_id' a -@[simp] -theorem lift_lift (a : Ordinal) : lift.{w} (lift.{v} a) = lift.{max v w} a := - inductionOn a fun _ _ _ => - Quotient.sound - ⟨(RelIso.preimage Equiv.ulift _).trans <| - (RelIso.preimage Equiv.ulift _).trans (RelIso.preimage Equiv.ulift _).symm⟩ - theorem lift_type_le {α : Type u} {β : Type v} {r s} [IsWellOrder α r] [IsWellOrder β s] : lift.{max v w} (type r) ≤ lift.{max u w} (type s) ↔ Nonempty (r ≼i s) := ⟨fun ⟨f⟩ => @@ -638,19 +631,50 @@ theorem lift_type_lt {α : Type u} {β : Type v} {r s} [IsWellOrder α r] [IsWel (InitialSeg.ofIso (RelIso.preimage Equiv.ulift s).symm)⟩⟩ @[simp] -theorem lift_le {a b : Ordinal} : lift.{u,v} a ≤ lift.{u,v} b ↔ a ≤ b := - inductionOn a fun α r _ => - inductionOn b fun β s _ => by - rw [← lift_umax] - exact lift_type_le.{_,_,u} +theorem lift_le {a b : Ordinal} : lift.{u, v} a ≤ lift.{u, v} b ↔ a ≤ b := + inductionOn₂ a b fun α r _ β s _ => by + rw [← lift_umax] + exact lift_type_le.{_,_,u} + +@[simp] +theorem lift_inj {a b : Ordinal} : lift.{u, v} a = lift.{u, v} b ↔ a = b := by + simp_rw [le_antisymm_iff, lift_le] + +@[simp] +theorem lift_lt {a b : Ordinal} : lift.{u, v} a < lift.{u, v} b ↔ a < b := by + simp_rw [lt_iff_le_not_le, lift_le] @[simp] -theorem lift_inj {a b : Ordinal} : lift.{u,v} a = lift.{u,v} b ↔ a = b := by - simp only [le_antisymm_iff, lift_le] +theorem lift_typein_top {r : α → α → Prop} {s : β → β → Prop} + [IsWellOrder α r] [IsWellOrder β s] (f : r ≺i s) : lift.{u} (typein s f.top) = lift (type r) := + f.subrelIso.ordinal_lift_type_eq + +/-- Initial segment version of the lift operation on ordinals, embedding `Ordinal.{u}` in +`Ordinal.{v}` as an initial segment when `u ≤ v`. -/ +def liftInitialSeg : Ordinal.{v} ≤i Ordinal.{max u v} := by + refine ⟨RelEmbedding.ofMonotone lift.{u} (by simp), + fun a b ↦ Ordinal.inductionOn₂ a b fun α r _ β s _ h ↦ ?_⟩ + rw [RelEmbedding.ofMonotone_coe, ← lift_id'.{max u v} (type s), + ← lift_umax.{v, u}, lift_type_lt] at h + obtain ⟨f⟩ := h + use typein r f.top + rw [RelEmbedding.ofMonotone_coe, ← lift_umax, lift_typein_top, lift_id'] + +@[deprecated liftInitialSeg (since := "2024-09-21")] +alias lift.initialSeg := liftInitialSeg + +@[simp] +theorem liftInitialSeg_coe : (liftInitialSeg.{v, u} : Ordinal → Ordinal) = lift.{v, u} := + rfl + +set_option linter.deprecated false in +@[deprecated liftInitialSeg_coe (since := "2024-09-21")] +theorem lift.initialSeg_coe : (lift.initialSeg.{v, u} : Ordinal → Ordinal) = lift.{v, u} := + rfl @[simp] -theorem lift_lt {a b : Ordinal} : lift.{u,v} a < lift.{u,v} b ↔ a < b := by - simp only [lt_iff_le_not_le, lift_le] +theorem lift_lift (a : Ordinal.{u}) : lift.{w} (lift.{v} a) = lift.{max v w} a := + (liftInitialSeg.trans liftInitialSeg).eq liftInitialSeg a @[simp] theorem lift_zero : lift 0 = 0 := @@ -661,59 +685,25 @@ theorem lift_one : lift 1 = 1 := type_eq_one_of_unique _ @[simp] -theorem lift_card (a) : Cardinal.lift.{u,v} (card a)= card (lift.{u,v} a) := +theorem lift_card (a) : Cardinal.lift.{u, v} (card a) = card (lift.{u} a) := inductionOn a fun _ _ _ => rfl -theorem lift_down' {a : Cardinal.{u}} {b : Ordinal.{max u v}} - (h : card.{max u v} b ≤ Cardinal.lift.{v,u} a) : ∃ a', lift.{v,u} a' = b := - let ⟨c, e⟩ := Cardinal.mem_range_of_le_lift h - Cardinal.inductionOn c - (fun α => - inductionOn b fun β s _ e' => by - rw [card_type, ← Cardinal.lift_id'.{max u v, u} #β, ← Cardinal.lift_umax.{u, v}, - lift_mk_eq.{u, max u v, max u v}] at e' - cases' e' with f - have g := RelIso.preimage f s - haveI := (g : f ⁻¹'o s ↪r s).isWellOrder - have := lift_type_eq.{u, max u v, max u v}.2 ⟨g⟩ - rw [lift_id, lift_umax.{u, v}] at this - exact ⟨_, this⟩) - e +theorem mem_range_lift_of_le {a : Ordinal.{u}} {b : Ordinal.{max u v}} (h : b ≤ lift.{v} a) : + b ∈ Set.range lift.{v} := + liftInitialSeg.mem_range_of_le h +@[deprecated mem_range_lift_of_le (since := "2024-10-07")] theorem lift_down {a : Ordinal.{u}} {b : Ordinal.{max u v}} (h : b ≤ lift.{v,u} a) : ∃ a', lift.{v,u} a' = b := - @lift_down' (card a) _ (by rw [lift_card]; exact card_le_card h) + mem_range_lift_of_le h theorem le_lift_iff {a : Ordinal.{u}} {b : Ordinal.{max u v}} : - b ≤ lift.{v,u} a ↔ ∃ a', lift.{v,u} a' = b ∧ a' ≤ a := - ⟨fun h => - let ⟨a', e⟩ := lift_down h - ⟨a', e, lift_le.1 <| e.symm ▸ h⟩, - fun ⟨_, e, h⟩ => e ▸ lift_le.2 h⟩ + b ≤ lift.{v} a ↔ ∃ a' ≤ a, lift.{v} a' = b := + liftInitialSeg.le_apply_iff theorem lt_lift_iff {a : Ordinal.{u}} {b : Ordinal.{max u v}} : - b < lift.{v,u} a ↔ ∃ a', lift.{v,u} a' = b ∧ a' < a := - ⟨fun h => - let ⟨a', e⟩ := lift_down (le_of_lt h) - ⟨a', e, lift_lt.1 <| e.symm ▸ h⟩, - fun ⟨_, e, h⟩ => e ▸ lift_lt.2 h⟩ - -/-- Initial segment version of the lift operation on ordinals, embedding `Ordinal.{u}` in -`Ordinal.{v}` as an initial segment when `u ≤ v`. -/ -def liftInitialSeg : Ordinal.{u} ≤i Ordinal.{max u v} := - ⟨⟨⟨lift.{v}, fun _ _ => lift_inj.1⟩, lift_lt⟩, fun _ _ h => lift_down (le_of_lt h)⟩ - -@[deprecated liftInitialSeg (since := "2024-09-21")] -alias lift.initialSeg := liftInitialSeg - -@[simp] -theorem liftInitialSeg_coe : (liftInitialSeg.{u, v} : Ordinal → Ordinal) = lift.{v, u} := - rfl - -set_option linter.deprecated false in -@[deprecated liftInitialSeg_coe (since := "2024-09-21")] -theorem lift.initialSeg_coe : (lift.initialSeg.{u, v} : Ordinal → Ordinal) = lift.{v, u} := - rfl + b < lift.{v} a ↔ ∃ a' < a, lift.{v} a' = b := + liftInitialSeg.lt_apply_iff /-! ### The first infinite ordinal ω -/ @@ -1026,7 +1016,7 @@ theorem univ_umax : univ.{u, max (u + 1) v} = univ.{u, v} := /-- Principal segment version of the lift operation on ordinals, embedding `Ordinal.{u}` in `Ordinal.{v}` as a principal segment when `u < v`. -/ def liftPrincipalSeg : Ordinal.{u} inductionOn b ?_; intro β s _ rw [univ, ← lift_umax]; constructor <;> intro h · cases' h with a e @@ -1220,14 +1210,14 @@ theorem ord_ofNat (n : ℕ) [n.AtLeastTwo] : ord (no_index (OfNat.ofNat n)) = Of theorem ord_aleph0 : ord.{u} ℵ₀ = ω := le_antisymm (ord_le.2 le_rfl) <| le_of_forall_lt fun o h => by - rcases Ordinal.lt_lift_iff.1 h with ⟨o, rfl, h'⟩ + rcases Ordinal.lt_lift_iff.1 h with ⟨o, h', rfl⟩ rw [lt_ord, ← lift_card, lift_lt_aleph0, ← typein_enum (· < ·) h'] exact lt_aleph0_iff_fintype.2 ⟨Set.fintypeLTNat _⟩ @[simp] theorem lift_ord (c) : Ordinal.lift.{u,v} (ord c) = ord (lift.{u,v} c) := by refine le_antisymm (le_of_forall_lt fun a ha => ?_) ?_ - · rcases Ordinal.lt_lift_iff.1 ha with ⟨a, rfl, _⟩ + · rcases Ordinal.lt_lift_iff.1 ha with ⟨a, _, rfl⟩ rwa [lt_ord, ← lift_card, lift_lt, ← lt_ord, ← Ordinal.lift_lt] · rw [ord_le, ← lift_card, card_ord] @@ -1419,6 +1409,16 @@ theorem card_eq_zero {o} : card o = 0 ↔ o = 0 := by theorem card_eq_one {o} : card o = 1 ↔ o = 1 := by simpa using card_eq_nat (n := 1) +theorem mem_range_lift_of_card_le {a : Cardinal.{u}} {b : Ordinal.{max u v}} + (h : card b ≤ Cardinal.lift.{v, u} a) : b ∈ Set.range lift.{v, u} := by + rw [card_le_iff, ← lift_succ, ← lift_ord] at h + exact mem_range_lift_of_le h.le + +@[deprecated mem_range_lift_of_card_le (since := "2024-10-07")] +theorem lift_down' {a : Cardinal.{u}} {b : Ordinal.{max u v}} + (h : card.{max u v} b ≤ Cardinal.lift.{v, u} a) : ∃ a', lift.{v, u} a' = b := + mem_range_lift_of_card_le h + -- See note [no_index around OfNat.ofNat] @[simp] theorem card_eq_ofNat {o} {n : ℕ} [n.AtLeastTwo] : diff --git a/Mathlib/SetTheory/ZFC/Rank.lean b/Mathlib/SetTheory/ZFC/Rank.lean index cdcacbd303f4e..9ab8f265dc4f9 100644 --- a/Mathlib/SetTheory/ZFC/Rank.lean +++ b/Mathlib/SetTheory/ZFC/Rank.lean @@ -116,7 +116,7 @@ theorem rank_eq_wfRank : lift.{u + 1, u} (rank x) = IsWellFounded.rank (α := PS simp_rw [← fun y : { y // y ∈ x } => ih y y.2] apply (le_of_forall_lt _).antisymm (Ordinal.iSup_le _) <;> intro h · rw [lt_lift_iff] - rintro ⟨o, rfl, h⟩ + rintro ⟨o, h, rfl⟩ simpa [Ordinal.lt_iSup] using lt_rank_iff.1 h · simpa using rank_lt_of_mem h.2 @@ -202,7 +202,7 @@ theorem rank_eq_wfRank : lift.{u + 1, u} (rank x) = IsWellFounded.rank (α := ZF simp_rw [← fun y : { y // y ∈ x } => ih y y.2] apply (le_of_forall_lt _).antisymm (Ordinal.iSup_le _) <;> intro h · rw [lt_lift_iff] - rintro ⟨o, rfl, h⟩ + rintro ⟨o, h, rfl⟩ simpa [Ordinal.lt_iSup] using lt_rank_iff.1 h · simpa using rank_lt_of_mem h.2 From 13acd64e16908213609da0f86a6ce0f02e52063c Mon Sep 17 00:00:00 2001 From: Pietro Monticone <38562595+pitmonticone@users.noreply.github.com> Date: Tue, 22 Oct 2024 18:07:14 +0000 Subject: [PATCH 296/505] feat(Data/Set/Function): add `BijOn.insert` and `BijOn.sdiff_singleton` (#17971) Co-authored-by: Daniel Weber Upstreamed from the [EquationalTheories](https://github.com/teorth/equational_theories) project. Co-authored-by: Daniel Weber --- Mathlib/Data/Set/Function.lean | 28 ++++++++++++++++++++++++++++ 1 file changed, 28 insertions(+) diff --git a/Mathlib/Data/Set/Function.lean b/Mathlib/Data/Set/Function.lean index 4943372c57ca0..1e7d3e37be5c1 100644 --- a/Mathlib/Data/Set/Function.lean +++ b/Mathlib/Data/Set/Function.lean @@ -947,6 +947,34 @@ theorem BijOn.subset_left {r : Set α} (hf : BijOn f s t) (hrs : r ⊆ s) : BijOn f r (f '' r) := (hf.injOn.mono hrs).bijOn_image +theorem BijOn.insert_iff (ha : a ∉ s) (hfa : f a ∉ t) : + BijOn f (insert a s) (insert (f a) t) ↔ BijOn f s t where + mp h := by + have := congrArg (· \ {f a}) (image_insert_eq ▸ h.image_eq) + simp only [mem_singleton_iff, insert_diff_of_mem] at this + rw [diff_singleton_eq_self hfa, diff_singleton_eq_self] at this + · exact ⟨by simp [← this, mapsTo'], h.injOn.mono (subset_insert ..), + by simp [← this, surjOn_image]⟩ + simp only [mem_image, not_exists, not_and] + intro x hx + rw [h.injOn.eq_iff (by simp [hx]) (by simp)] + exact ha ∘ (· ▸ hx) + mpr h := by + repeat rw [insert_eq] + refine (bijOn_singleton.mpr rfl).union h ?_ + simp only [singleton_union, injOn_insert fun x ↦ (hfa (h.mapsTo x)), h.injOn, mem_image, + not_exists, not_and, true_and] + exact fun _ hx h₂ ↦ hfa (h₂ ▸ h.mapsTo hx) + +theorem BijOn.insert (h₁ : BijOn f s t) (h₂ : f a ∉ t) : + BijOn f (insert a s) (insert (f a) t) := + (insert_iff (h₂ <| h₁.mapsTo ·) h₂).mpr h₁ + +theorem BijOn.sdiff_singleton (h₁ : BijOn f s t) (h₂ : a ∈ s) : + BijOn f (s \ {a}) (t \ {f a}) := by + convert h₁.subset_left diff_subset + simp [h₁.injOn.image_diff, h₁.image_eq, h₂, inter_eq_self_of_subset_right] + end bijOn /-! ### left inverse -/ From 5ddbb5a9a9cf248348cc06c546608e8a5cb02f5b Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Tue, 22 Oct 2024 18:07:15 +0000 Subject: [PATCH 297/505] chore: delete "`simp` can prove this" porting notes (#18080) This PR checks all porting notes of the form "`simp` can prove this" and if this is indeed the case (and the proof appears to make sense), removes the associated porting note. This is a somewhat random selection of folders that only contained a few porting notes. --- .../Additive/AP/Three/Behrend.lean | 1 - .../Combinatorics/SimpleGraph/AdjMatrix.lean | 1 - Mathlib/Combinatorics/SimpleGraph/Prod.lean | 2 -- Mathlib/Control/ULiftable.lean | 2 -- .../FieldTheory/IntermediateField/Basic.lean | 2 +- .../Euclidean/Angle/Oriented/Basic.lean | 4 --- Mathlib/Geometry/Euclidean/Basic.lean | 1 - Mathlib/Geometry/Manifold/Diffeomorph.lean | 2 -- Mathlib/InformationTheory/Hamming.lean | 7 ---- Mathlib/Logic/Equiv/Defs.lean | 4 --- Mathlib/Logic/Equiv/Set.lean | 1 - .../ProbabilityMassFunction/Monad.lean | 1 - Mathlib/SetTheory/Cardinal/Arithmetic.lean | 3 -- Mathlib/SetTheory/Cardinal/Basic.lean | 34 ------------------- Mathlib/SetTheory/Ordinal/Arithmetic.lean | 1 - Mathlib/SetTheory/Ordinal/Notation.lean | 1 - 16 files changed, 1 insertion(+), 66 deletions(-) diff --git a/Mathlib/Combinatorics/Additive/AP/Three/Behrend.lean b/Mathlib/Combinatorics/Additive/AP/Three/Behrend.lean index 54cfd59538a92..b0e0d2360ce16 100644 --- a/Mathlib/Combinatorics/Additive/AP/Three/Behrend.lean +++ b/Mathlib/Combinatorics/Additive/AP/Three/Behrend.lean @@ -127,7 +127,6 @@ def map (d : ℕ) : (Fin n → ℕ) →+ ℕ where map_zero' := by simp_rw [Pi.zero_apply, zero_mul, sum_const_zero] map_add' a b := by simp_rw [Pi.add_apply, add_mul, sum_add_distrib] --- @[simp] -- Porting note (#10618): simp can prove this theorem map_zero (d : ℕ) (a : Fin 0 → ℕ) : map d a = 0 := by simp [map] theorem map_succ (a : Fin (n + 1) → ℕ) : diff --git a/Mathlib/Combinatorics/SimpleGraph/AdjMatrix.lean b/Mathlib/Combinatorics/SimpleGraph/AdjMatrix.lean index 338e9167ee4a1..a1c4a1e4dc682 100644 --- a/Mathlib/Combinatorics/SimpleGraph/AdjMatrix.lean +++ b/Mathlib/Combinatorics/SimpleGraph/AdjMatrix.lean @@ -231,7 +231,6 @@ theorem adjMatrix_mul_self_apply_self [NonAssocSemiring α] (i : V) : variable {G} --- @[simp] -- Porting note (#10618): simp can prove this theorem adjMatrix_mulVec_const_apply [NonAssocSemiring α] {a : α} {v : V} : (G.adjMatrix α *ᵥ Function.const _ a) v = G.degree v * a := by simp diff --git a/Mathlib/Combinatorics/SimpleGraph/Prod.lean b/Mathlib/Combinatorics/SimpleGraph/Prod.lean index df7d60cc174f7..e094ff9729ec2 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Prod.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Prod.lean @@ -50,12 +50,10 @@ theorem boxProd_adj {x y : α × β} : (G □ H).Adj x y ↔ G.Adj x.1 y.1 ∧ x.2 = y.2 ∨ H.Adj x.2 y.2 ∧ x.1 = y.1 := Iff.rfl ---@[simp] Porting note (#10618): `simp` can prove theorem boxProd_adj_left {a₁ : α} {b : β} {a₂ : α} : (G □ H).Adj (a₁, b) (a₂, b) ↔ G.Adj a₁ a₂ := by simp only [boxProd_adj, and_true, SimpleGraph.irrefl, false_and, or_false] ---@[simp] Porting note (#10618): `simp` can prove theorem boxProd_adj_right {a : α} {b₁ b₂ : β} : (G □ H).Adj (a, b₁) (a, b₂) ↔ H.Adj b₁ b₂ := by simp only [boxProd_adj, SimpleGraph.irrefl, false_and, and_true, false_or] diff --git a/Mathlib/Control/ULiftable.lean b/Mathlib/Control/ULiftable.lean index 50eae9d2f94a5..23f4e9733b8a4 100644 --- a/Mathlib/Control/ULiftable.lean +++ b/Mathlib/Control/ULiftable.lean @@ -93,12 +93,10 @@ def downMap {F : Type max u₀ v₀ → Type u₁} {G : Type u₀ → Type v₁} [Functor F] {α β} (f : α → β) (x : F α) : G β := down (Functor.map (ULift.up.{v₀} ∘ f) x : F (ULift β)) --- @[simp] -- Porting note (#10618): simp can prove this theorem up_down {f : Type u₀ → Type u₁} {g : Type max u₀ v₀ → Type v₁} [ULiftable f g] {α} (x : g (ULift.{v₀} α)) : up (down x : f α) = x := (ULiftable.congr Equiv.ulift.symm).right_inv _ --- @[simp] -- Porting note (#10618): simp can prove this theorem down_up {f : Type u₀ → Type u₁} {g : Type max u₀ v₀ → Type v₁} [ULiftable f g] {α} (x : f α) : down (up x : g (ULift.{v₀} α)) = x := (ULiftable.congr Equiv.ulift.symm).left_inv _ diff --git a/Mathlib/FieldTheory/IntermediateField/Basic.lean b/Mathlib/FieldTheory/IntermediateField/Basic.lean index 54dbf4d6b374d..bd1efcb1176bb 100644 --- a/Mathlib/FieldTheory/IntermediateField/Basic.lean +++ b/Mathlib/FieldTheory/IntermediateField/Basic.lean @@ -295,7 +295,7 @@ theorem coe_sum {ι : Type*} [Fintype ι] (f : ι → S) : (↑(∑ i, f i) : L) · simp · rw [Finset.sum_insert hi, AddMemClass.coe_add, H, Finset.sum_insert hi] -@[norm_cast] --Porting note (#10618): `simp` can prove it +@[norm_cast] theorem coe_prod {ι : Type*} [Fintype ι] (f : ι → S) : (↑(∏ i, f i) : L) = ∏ i, (f i : L) := by classical induction' (Finset.univ : Finset ι) using Finset.induction_on with i s hi H diff --git a/Mathlib/Geometry/Euclidean/Angle/Oriented/Basic.lean b/Mathlib/Geometry/Euclidean/Angle/Oriented/Basic.lean index 0ad26516219b6..c36abe59e9baf 100644 --- a/Mathlib/Geometry/Euclidean/Angle/Oriented/Basic.lean +++ b/Mathlib/Geometry/Euclidean/Angle/Oriented/Basic.lean @@ -223,12 +223,10 @@ theorem oangle_neg_self_right {x : V} (hx : x ≠ 0) : o.oangle x (-x) = π := b simp [oangle_neg_right, hx] /-- Twice the angle between the negation of a vector and that vector is 0. -/ --- @[simp] -- Porting note (#10618): simp can prove this theorem two_zsmul_oangle_neg_self_left (x : V) : (2 : ℤ) • o.oangle (-x) x = 0 := by by_cases hx : x = 0 <;> simp [hx] /-- Twice the angle between a vector and its negation is 0. -/ --- @[simp] -- Porting note (#10618): simp can prove this theorem two_zsmul_oangle_neg_self_right (x : V) : (2 : ℤ) • o.oangle x (-x) = 0 := by by_cases hx : x = 0 <;> simp [hx] @@ -901,7 +899,6 @@ theorem oangle_sign_sub_left_swap (x y : V) : (o.oangle (x - y) x).sign = (o.oan /-- The sign of the angle between a vector, and a linear combination of that vector with a second vector, is the sign of the factor by which the second vector is multiplied in that combination multiplied by the sign of the angle between the two vectors. -/ --- @[simp] -- Porting note (#10618): simp can prove this theorem oangle_sign_smul_add_smul_right (x y : V) (r₁ r₂ : ℝ) : (o.oangle x (r₁ • x + r₂ • y)).sign = SignType.sign r₂ * (o.oangle x y).sign := by rw [← o.oangle_sign_smul_add_right x (r₁ • x + r₂ • y) (-r₁)] @@ -910,7 +907,6 @@ theorem oangle_sign_smul_add_smul_right (x y : V) (r₁ r₂ : ℝ) : /-- The sign of the angle between a linear combination of two vectors and the second vector is the sign of the factor by which the first vector is multiplied in that combination multiplied by the sign of the angle between the two vectors. -/ --- @[simp] -- Porting note (#10618): simp can prove this theorem oangle_sign_smul_add_smul_left (x y : V) (r₁ r₂ : ℝ) : (o.oangle (r₁ • x + r₂ • y) y).sign = SignType.sign r₁ * (o.oangle x y).sign := by simp_rw [o.oangle_rev y, Real.Angle.sign_neg, add_comm (r₁ • x), oangle_sign_smul_add_smul_right, diff --git a/Mathlib/Geometry/Euclidean/Basic.lean b/Mathlib/Geometry/Euclidean/Basic.lean index b9807b923e2cb..d39b79aad2703 100644 --- a/Mathlib/Geometry/Euclidean/Basic.lean +++ b/Mathlib/Geometry/Euclidean/Basic.lean @@ -353,7 +353,6 @@ theorem orthogonalProjection_mem_subspace_eq_self {s : AffineSubspace ℝ P} [No exact p.2 /-- Orthogonal projection is idempotent. -/ --- @[simp] -- Porting note (#10618): simp can prove this theorem orthogonalProjection_orthogonalProjection (s : AffineSubspace ℝ P) [Nonempty s] [HasOrthogonalProjection s.direction] (p : P) : orthogonalProjection s (orthogonalProjection s p) = orthogonalProjection s p := by diff --git a/Mathlib/Geometry/Manifold/Diffeomorph.lean b/Mathlib/Geometry/Manifold/Diffeomorph.lean index 425037472ff7a..fa71b1fc3e365 100644 --- a/Mathlib/Geometry/Manifold/Diffeomorph.lean +++ b/Mathlib/Geometry/Manifold/Diffeomorph.lean @@ -529,7 +529,6 @@ theorem contMDiff_transDiffeomorph_right {f : M' → M} : ContMDiff I' (I.transDiffeomorph e) n f ↔ ContMDiff I' I n f := (toTransDiffeomorph I M e).contMDiff_diffeomorph_comp_iff le_top --- Porting note (#10618): was `@[simp]` but now `simp` can prove it theorem smooth_transDiffeomorph_right {f : M' → M} : Smooth I' (I.transDiffeomorph e) f ↔ Smooth I' I f := contMDiff_transDiffeomorph_right e @@ -554,7 +553,6 @@ theorem contMDiff_transDiffeomorph_left {f : M → M'} : ContMDiff (I.transDiffeomorph e) I' n f ↔ ContMDiff I I' n f := ((toTransDiffeomorph I M e).contMDiff_comp_diffeomorph_iff le_top).symm --- Porting note (#10618): was `@[simp]` but now `simp` can prove it theorem smooth_transDiffeomorph_left {f : M → M'} : Smooth (I.transDiffeomorph e) I' f ↔ Smooth I I' f := e.contMDiff_transDiffeomorph_left diff --git a/Mathlib/InformationTheory/Hamming.lean b/Mathlib/InformationTheory/Hamming.lean index 63b352144a998..1cd0c586d3812 100644 --- a/Mathlib/InformationTheory/Hamming.lean +++ b/Mathlib/InformationTheory/Hamming.lean @@ -103,7 +103,6 @@ theorem hammingDist_ne_zero {x y : ∀ i, β i} : hammingDist x y ≠ 0 ↔ x theorem hammingDist_pos {x y : ∀ i, β i} : 0 < hammingDist x y ↔ x ≠ y := by rw [← hammingDist_ne_zero, iff_not_comm, not_lt, Nat.le_zero] --- @[simp] -- Porting note (#10618): simp can prove this theorem hammingDist_lt_one {x y : ∀ i, β i} : hammingDist x y < 1 ↔ x = y := by rw [Nat.lt_one_iff, hammingDist_eq_zero] @@ -147,7 +146,6 @@ theorem hammingDist_zero_left : hammingDist (0 : ∀ i, β i) = hammingNorm := funext fun x => by rw [hammingDist_comm, hammingDist_zero_right] /-- Corresponds to `norm_nonneg`. -/ --- @[simp] -- Porting note (#10618): simp can prove this theorem hammingNorm_nonneg {x : ∀ i, β i} : 0 ≤ hammingNorm x := zero_le _ @@ -170,7 +168,6 @@ theorem hammingNorm_ne_zero_iff {x : ∀ i, β i} : hammingNorm x ≠ 0 ↔ x theorem hammingNorm_pos_iff {x : ∀ i, β i} : 0 < hammingNorm x ↔ x ≠ 0 := hammingDist_pos --- @[simp] -- Porting note (#10618): simp can prove this theorem hammingNorm_lt_one {x : ∀ i, β i} : hammingNorm x < 1 ↔ x = 0 := hammingDist_lt_one @@ -288,13 +285,9 @@ theorem toHamming_ofHamming (x : Hamming β) : toHamming (ofHamming x) = x := theorem ofHamming_toHamming (x : ∀ i, β i) : ofHamming (toHamming x) = x := rfl ---@[simp] -- Porting note (#10618): removing `simp`, `simp` can prove it --- and `dsimp` cannot use `Iff.rfl` theorem toHamming_inj {x y : ∀ i, β i} : toHamming x = toHamming y ↔ x = y := Iff.rfl ---@[simp] -- Porting note (#10618): removing `simp`, `simp` can prove it --- and `dsimp` cannot use `Iff.rfl` theorem ofHamming_inj {x y : Hamming β} : ofHamming x = ofHamming y ↔ x = y := Iff.rfl diff --git a/Mathlib/Logic/Equiv/Defs.lean b/Mathlib/Logic/Equiv/Defs.lean index a8fbb1a0ed801..0cda134871dce 100644 --- a/Mathlib/Logic/Equiv/Defs.lean +++ b/Mathlib/Logic/Equiv/Defs.lean @@ -591,13 +591,11 @@ def psigmaCongrRight {β₁ β₂ : α → Sort*} (F : ∀ a, β₁ a ≃ β₂ left_inv | ⟨a, b⟩ => congr_arg (PSigma.mk a) <| symm_apply_apply (F a) b right_inv | ⟨a, b⟩ => congr_arg (PSigma.mk a) <| apply_symm_apply (F a) b --- Porting note (#10618): simp can now simplify the LHS, so I have removed `@[simp]` theorem psigmaCongrRight_trans {α} {β₁ β₂ β₃ : α → Sort*} (F : ∀ a, β₁ a ≃ β₂ a) (G : ∀ a, β₂ a ≃ β₃ a) : (psigmaCongrRight F).trans (psigmaCongrRight G) = psigmaCongrRight fun a => (F a).trans (G a) := rfl --- Porting note (#10618): simp can now simplify the LHS, so I have removed `@[simp]` theorem psigmaCongrRight_symm {α} {β₁ β₂ : α → Sort*} (F : ∀ a, β₁ a ≃ β₂ a) : (psigmaCongrRight F).symm = psigmaCongrRight fun a => (F a).symm := rfl @@ -614,13 +612,11 @@ def sigmaCongrRight {α} {β₁ β₂ : α → Type*} (F : ∀ a, β₁ a ≃ β left_inv | ⟨a, b⟩ => congr_arg (Sigma.mk a) <| symm_apply_apply (F a) b right_inv | ⟨a, b⟩ => congr_arg (Sigma.mk a) <| apply_symm_apply (F a) b --- Porting note (#10618): simp can now simplify the LHS, so I have removed `@[simp]` theorem sigmaCongrRight_trans {α} {β₁ β₂ β₃ : α → Type*} (F : ∀ a, β₁ a ≃ β₂ a) (G : ∀ a, β₂ a ≃ β₃ a) : (sigmaCongrRight F).trans (sigmaCongrRight G) = sigmaCongrRight fun a => (F a).trans (G a) := rfl --- Porting note (#10618): simp can now simplify the LHS, so I have removed `@[simp]` theorem sigmaCongrRight_symm {α} {β₁ β₂ : α → Type*} (F : ∀ a, β₁ a ≃ β₂ a) : (sigmaCongrRight F).symm = sigmaCongrRight fun a => (F a).symm := rfl diff --git a/Mathlib/Logic/Equiv/Set.lean b/Mathlib/Logic/Equiv/Set.lean index a734cb9993db0..641faf211dc32 100644 --- a/Mathlib/Logic/Equiv/Set.lean +++ b/Mathlib/Logic/Equiv/Set.lean @@ -104,7 +104,6 @@ theorem preimage_symm_preimage {α β} (e : α ≃ β) (s : Set α) : e ⁻¹' ( theorem preimage_subset {α β} (e : α ≃ β) (s t : Set β) : e ⁻¹' s ⊆ e ⁻¹' t ↔ s ⊆ t := e.surjective.preimage_subset_preimage_iff --- Porting note (#10618): removed `simp` attribute. `simp` can prove it. theorem image_subset {α β} (e : α ≃ β) (s t : Set α) : e '' s ⊆ e '' t ↔ s ⊆ t := image_subset_image_iff e.injective diff --git a/Mathlib/Probability/ProbabilityMassFunction/Monad.lean b/Mathlib/Probability/ProbabilityMassFunction/Monad.lean index 3e0a147acaeca..acf07b214fb38 100644 --- a/Mathlib/Probability/ProbabilityMassFunction/Monad.lean +++ b/Mathlib/Probability/ProbabilityMassFunction/Monad.lean @@ -48,7 +48,6 @@ theorem support_pure : (pure a).support = {a} := theorem mem_support_pure_iff : a' ∈ (pure a).support ↔ a' = a := by simp --- @[simp] -- Porting note (#10618): simp can prove this theorem pure_apply_self : pure a a = 1 := if_pos rfl diff --git a/Mathlib/SetTheory/Cardinal/Arithmetic.lean b/Mathlib/SetTheory/Cardinal/Arithmetic.lean index b942a95f09daa..a002340504602 100644 --- a/Mathlib/SetTheory/Cardinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Cardinal/Arithmetic.lean @@ -111,11 +111,9 @@ theorem aleph0_mul_eq {a : Cardinal} (ha : ℵ₀ ≤ a) : ℵ₀ * a = a := theorem mul_aleph0_eq {a : Cardinal} (ha : ℵ₀ ≤ a) : a * ℵ₀ = a := (mul_eq_max ha le_rfl).trans (max_eq_left ha) --- Porting note (#10618): removed `simp`, `simp` can prove it theorem aleph0_mul_mk_eq {α : Type*} [Infinite α] : ℵ₀ * #α = #α := aleph0_mul_eq (aleph0_le_mk α) --- Porting note (#10618): removed `simp`, `simp` can prove it theorem mk_mul_aleph0_eq {α : Type*} [Infinite α] : #α * ℵ₀ = #α := mul_aleph0_eq (aleph0_le_mk α) @@ -325,7 +323,6 @@ theorem nat_add_eq {a : Cardinal} (n : ℕ) (ha : ℵ₀ ≤ a) : n + a = a := b theorem add_one_eq {a : Cardinal} (ha : ℵ₀ ≤ a) : a + 1 = a := add_one_of_aleph0_le ha --- Porting note (#10618): removed `simp`, `simp` can prove it theorem mk_add_one_eq {α : Type*} [Infinite α] : #α + 1 = #α := add_one_eq (aleph0_le_mk α) diff --git a/Mathlib/SetTheory/Cardinal/Basic.lean b/Mathlib/SetTheory/Cardinal/Basic.lean index f9e31b0db1b7e..006f281adb63c 100644 --- a/Mathlib/SetTheory/Cardinal/Basic.lean +++ b/Mathlib/SetTheory/Cardinal/Basic.lean @@ -259,8 +259,6 @@ theorem lift_id (a : Cardinal) : lift.{u, u} a = a := lift_id'.{u, u} a /-- A cardinal lifted to the zero universe equals itself. -/ --- porting note (#10618): simp can prove this --- @[simp] theorem lift_uzero (a : Cardinal.{u}) : lift.{0} a = a := lift_id'.{0, u} a @@ -541,12 +539,8 @@ instance commSemiring : CommSemiring Cardinal.{u} where theorem one_power {a : Cardinal} : (1 : Cardinal) ^ a = 1 := inductionOn a fun _ => mk_eq_one _ --- porting note (#10618): simp can prove this --- @[simp] theorem mk_bool : #Bool = 2 := by simp --- porting note (#10618): simp can prove this --- @[simp] theorem mk_Prop : #Prop = 2 := by simp @[simp] @@ -1275,8 +1269,6 @@ theorem lift_eq_aleph0 {c : Cardinal.{u}} : lift.{v} c = ℵ₀ ↔ c = ℵ₀ : /-! ### Properties about the cast from `ℕ` -/ --- porting note (#10618): simp can prove this --- @[simp] theorem mk_fin (n : ℕ) : #(Fin n) = n := by simp @[simp] @@ -1389,21 +1381,15 @@ theorem mk_finset_of_fintype [Fintype α] : #(Finset α) = 2 ^ Fintype.card α : theorem card_le_of_finset {α} (s : Finset α) : (s.card : Cardinal) ≤ #α := @mk_coe_finset _ s ▸ mk_set_le _ --- Porting note (#11119): was `simp`. LHS is not normal form. --- @[simp, norm_cast] @[norm_cast] theorem natCast_pow {m n : ℕ} : (↑(m ^ n) : Cardinal) = (↑m : Cardinal) ^ (↑n : Cardinal) := by induction n <;> simp [pow_succ, power_add, *, Pow.pow] --- porting note (#10618): simp can prove this --- @[simp, norm_cast] @[norm_cast] theorem natCast_le {m n : ℕ} : (m : Cardinal) ≤ n ↔ m ≤ n := by rw [← lift_mk_fin, ← lift_mk_fin, lift_le, le_def, Function.Embedding.nonempty_iff_card_le, Fintype.card_fin, Fintype.card_fin] --- porting note (#10618): simp can prove this --- @[simp, norm_cast] @[norm_cast] theorem natCast_lt {m n : ℕ} : (m : Cardinal) < n ↔ m < n := by rw [lt_iff_le_not_le, ← not_le] @@ -1569,8 +1555,6 @@ theorem lt_aleph0_iff_fintype {α : Type u} : #α < ℵ₀ ↔ Nonempty (Fintype theorem lt_aleph0_of_finite (α : Type u) [Finite α] : #α < ℵ₀ := lt_aleph0_iff_finite.2 ‹_› --- porting note (#10618): simp can prove this --- @[simp] theorem lt_aleph0_iff_set_finite {S : Set α} : #S < ℵ₀ ↔ S.Finite := lt_aleph0_iff_finite.trans finite_coe_iff @@ -1587,8 +1571,6 @@ theorem mk_le_aleph0_iff : #α ≤ ℵ₀ ↔ Countable α := by theorem mk_le_aleph0 [Countable α] : #α ≤ ℵ₀ := mk_le_aleph0_iff.mpr ‹_› --- porting note (#10618): simp can prove this --- @[simp] theorem le_aleph0_iff_set_countable {s : Set α} : #s ≤ ℵ₀ ↔ s.Countable := mk_le_aleph0_iff alias ⟨_, _root_.Set.Countable.le_aleph0⟩ := le_aleph0_iff_set_countable @@ -1691,8 +1673,6 @@ theorem denumerable_iff {α : Type u} : Nonempty (Denumerable α) ↔ #α = ℵ cases' Quotient.exact h with f exact ⟨Denumerable.mk' <| f.trans Equiv.ulift⟩⟩ --- porting note (#10618): simp can prove this --- @[simp] theorem mk_denumerable (α : Type u) [Denumerable α] : #α = ℵ₀ := denumerable_iff.1 ⟨‹_›⟩ @@ -1760,36 +1740,24 @@ theorem mk_pNat : #ℕ+ = ℵ₀ := /-! ### Cardinalities of basic sets and types -/ --- porting note (#10618): simp can prove this --- @[simp] theorem mk_empty : #Empty = 0 := mk_eq_zero _ --- porting note (#10618): simp can prove this --- @[simp] theorem mk_pempty : #PEmpty = 0 := mk_eq_zero _ --- porting note (#10618): simp can prove this --- @[simp] theorem mk_punit : #PUnit = 1 := mk_eq_one PUnit theorem mk_unit : #Unit = 1 := mk_punit --- porting note (#10618): simp can prove this --- @[simp] theorem mk_singleton {α : Type u} (x : α) : #({x} : Set α) = 1 := mk_eq_one _ --- porting note (#10618): simp can prove this --- @[simp] theorem mk_plift_true : #(PLift True) = 1 := mk_eq_one _ --- porting note (#10618): simp can prove this --- @[simp] theorem mk_plift_false : #(PLift False) = 0 := mk_eq_zero _ @@ -1812,8 +1780,6 @@ theorem mk_subtype_le_of_subset {α : Type u} {p q : α → Prop} (h : ∀ ⦃x #(Subtype p) ≤ #(Subtype q) := ⟨Embedding.subtypeMap (Embedding.refl α) h⟩ --- porting note (#10618): simp can prove this --- @[simp] theorem mk_emptyCollection (α : Type u) : #(∅ : Set α) = 0 := mk_eq_zero _ diff --git a/Mathlib/SetTheory/Ordinal/Arithmetic.lean b/Mathlib/SetTheory/Ordinal/Arithmetic.lean index 9725bed77ccd6..15e30c17e0a6f 100644 --- a/Mathlib/SetTheory/Ordinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Ordinal/Arithmetic.lean @@ -578,7 +578,6 @@ theorem isLimit_sub {a b} (l : IsLimit a) (h : b < a) : IsLimit (a - b) := @[deprecated isLimit_sub (since := "2024-10-11")] alias sub_isLimit := isLimit_sub --- @[simp] -- Porting note (#10618): simp can prove this theorem one_add_omega0 : 1 + ω = ω := by refine le_antisymm ?_ (le_add_left _ _) rw [omega0, ← lift_one.{0}, ← lift_add, lift_le, ← type_unit, ← type_sum_lex] diff --git a/Mathlib/SetTheory/Ordinal/Notation.lean b/Mathlib/SetTheory/Ordinal/Notation.lean index bab9ce4002c72..599622b062840 100644 --- a/Mathlib/SetTheory/Ordinal/Notation.lean +++ b/Mathlib/SetTheory/Ordinal/Notation.lean @@ -134,7 +134,6 @@ theorem ofNat_one : ofNat 1 = 1 := @[simp] theorem repr_ofNat (n : ℕ) : repr (ofNat n) = n := by cases n <;> simp --- @[simp] -- Porting note (#10618): simp can prove this theorem repr_one : repr (ofNat 1) = (1 : ℕ) := repr_ofNat 1 theorem omega0_le_oadd (e n a) : ω ^ repr e ≤ repr (oadd e n a) := by From 7d93d007291e4b6188a186c5e8660cf62c162815 Mon Sep 17 00:00:00 2001 From: Jakob von Raumer Date: Tue, 22 Oct 2024 18:49:13 +0000 Subject: [PATCH 298/505] feat(CategoryTheory/Abelian): AB5 implies AB4 (#17406) An AB4 category is a category in which taking coproducts is exact, an AB5 category is one where taking filtered colimits is exact. This shows that any Abelian catgory which is AB5 is also AB4. --- .../Abelian/GrothendieckAxioms.lean | 41 +++++++++++++++++-- .../Limits/Constructions/Filtered.lean | 12 ++++++ .../Limits/Preserves/FunctorCategory.lean | 32 ++++++++++----- 3 files changed, 72 insertions(+), 13 deletions(-) diff --git a/Mathlib/CategoryTheory/Abelian/GrothendieckAxioms.lean b/Mathlib/CategoryTheory/Abelian/GrothendieckAxioms.lean index ff25be4567c6f..84eda8c006ac2 100644 --- a/Mathlib/CategoryTheory/Abelian/GrothendieckAxioms.lean +++ b/Mathlib/CategoryTheory/Abelian/GrothendieckAxioms.lean @@ -4,8 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Isaac Hernando, Coleton Kotch, Adam Topaz -/ -import Mathlib.CategoryTheory.Limits.Filtered -import Mathlib.CategoryTheory.Limits.Preserves.Finite +import Mathlib.CategoryTheory.Limits.Constructions.Filtered +import Mathlib.CategoryTheory.Limits.Shapes.Biproducts +import Mathlib.CategoryTheory.Limits.Preserves.FunctorCategory /-! @@ -20,6 +21,10 @@ basic facts about them. - `AB5` -- an abelian category satisfies `AB5` provided that filtered colimits are exact. - The duals of the above definitions, called `AB4Star` and `AB5Star`. +## Theorems + +- The implication from `AB5` to `AB4` is established in `AB4.ofAB5`. + ## Remarks For `AB4` and `AB5`, we only require left exactness as right exactness is automatic. @@ -42,10 +47,12 @@ namespace CategoryTheory open Limits -universe v v' u u' +universe v v' u u' w variable (C : Type u) [Category.{v} C] +attribute [local instance] hasCoproducts_of_finite_and_filtered + /-- A category `C` which has coproducts is said to have `AB4` provided that coproducts are exact. @@ -90,4 +97,32 @@ class AB5Star [HasCofilteredLimits C] where attribute [instance] AB5Star.preservesFiniteColimits +noncomputable section + +open CoproductsFromFiniteFiltered + +variable {α : Type w} +variable [HasZeroMorphisms C] [HasFiniteBiproducts C] [HasFiniteLimits C] + +instance preservesFiniteLimitsLiftToFinset : PreservesFiniteLimits (liftToFinset C α) := + preservesFiniteLimitsOfEvaluation _ fun I => + letI : PreservesFiniteLimits (colim (J := Discrete I) (C := C)) := + preservesFiniteLimitsOfNatIso HasBiproductsOfShape.colimIsoLim.symm + letI : PreservesFiniteLimits ((whiskeringLeft (Discrete I) (Discrete α) C).obj + (Discrete.functor fun x ↦ ↑x)) := + ⟨fun J _ _ => whiskeringLeftPreservesLimitsOfShape J _⟩ + letI : PreservesFiniteLimits ((whiskeringLeft (Discrete I) (Discrete α) C).obj + (Discrete.functor (·.val)) ⋙ colim) := + compPreservesFiniteLimits _ _ + preservesFiniteLimitsOfNatIso (liftToFinsetEvaluationIso I).symm + +/-- A category with finite biproducts and finite limits is AB4 if it is AB5. -/ +def AB4.ofAB5 [HasFiniteCoproducts C] [HasFilteredColimits C] [AB5 C] : AB4 C where + preservesFiniteLimits J := + letI : PreservesFiniteLimits (liftToFinset C J ⋙ colim) := + compPreservesFiniteLimits _ _ + preservesFiniteLimitsOfNatIso (liftToFinsetColimIso) + +end + end CategoryTheory diff --git a/Mathlib/CategoryTheory/Limits/Constructions/Filtered.lean b/Mathlib/CategoryTheory/Limits/Constructions/Filtered.lean index 423e5cc93ea1b..e7e1a5ed68f8b 100644 --- a/Mathlib/CategoryTheory/Limits/Constructions/Filtered.lean +++ b/Mathlib/CategoryTheory/Limits/Constructions/Filtered.lean @@ -104,6 +104,8 @@ theorem has_limits_of_finite_and_cofiltered [HasFiniteLimits C] namespace CoproductsFromFiniteFiltered +section + variable [HasFiniteCoproducts C] [HasFilteredColimitsOfSize.{w, w} C] attribute [local instance] hasCoproducts_of_finite_and_filtered @@ -130,6 +132,16 @@ def liftToFinsetColimIso : liftToFinset C α ⋙ colim ≅ colim := Discrete.natTrans_app, liftToFinsetColimIso_aux, liftToFinsetColimIso_aux_assoc, ι_colimMap]) +end + +/-- `liftToFinset`, when composed with the evaluation functor, results in the whiskering composed +with `colim`. -/ +def liftToFinsetEvaluationIso [HasFiniteCoproducts C] (I : Finset (Discrete α)) : + liftToFinset C α ⋙ (evaluation _ _).obj I ≅ + (whiskeringLeft _ _ _).obj (Discrete.functor (·.val)) ⋙ colim (J := Discrete I) := + NatIso.ofComponents (fun _ => HasColimit.isoOfNatIso (Discrete.natIso fun _ => Iso.refl _)) + fun _ => by dsimp; ext; simp + end CoproductsFromFiniteFiltered end CategoryTheory.Limits diff --git a/Mathlib/CategoryTheory/Limits/Preserves/FunctorCategory.lean b/Mathlib/CategoryTheory/Limits/Preserves/FunctorCategory.lean index 0d33dbbf72e5a..23b061baf7be2 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/FunctorCategory.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/FunctorCategory.lean @@ -29,7 +29,7 @@ https://ncatlab.org/nlab/show/commutativity+of+limits+and+colimits#preservation_ -/ -universe w w' v₁ v₂ u u₂ +universe w w' v v₁ v₂ v₃ u u₁ u₂ u₃ noncomputable section @@ -37,6 +37,8 @@ namespace CategoryTheory open Category Limits +section + variable {C : Type u} [Category.{v₁} C] variable {D : Type u₂} [Category.{u} D] variable {E : Type u} [Category.{v₂} E] @@ -71,15 +73,25 @@ def FunctorCategory.prodPreservesColimits [HasBinaryProducts D] [HasColimits D] · intro G G' apply prodComparison_natural ((evaluation C D).obj k) (𝟙 F) } ) } -instance whiskeringLeftPreservesLimits [HasLimits D] (F : C ⥤ E) : - PreservesLimits ((whiskeringLeft C E D).obj F) := - ⟨fun {J} [hJ : Category J] => - ⟨fun {K} => - ⟨fun c {hc} => by - apply evaluationJointlyReflectsLimits - intro Y - change IsLimit (((evaluation E D).obj (F.obj Y)).mapCone c) - exact PreservesLimit.preserves hc⟩⟩⟩ +end + +variable {C : Type u₁} [Category.{v₁} C] +variable {D : Type u₂} [Category.{v₂} D] +variable {E : Type u₃} [Category.{v₃} E] + +instance whiskeringLeftPreservesLimitsOfShape (J : Type u) [Category.{v} J] + [HasLimitsOfShape J D] (F : C ⥤ E) : + PreservesLimitsOfShape J ((whiskeringLeft C E D).obj F) := + ⟨fun {K} => + ⟨fun c {hc} => by + apply evaluationJointlyReflectsLimits + intro Y + change IsLimit (((evaluation E D).obj (F.obj Y)).mapCone c) + exact PreservesLimit.preserves hc⟩⟩ + +instance whiskeringLeftPreservesLimits [HasLimitsOfSize.{w} D] (F : C ⥤ E) : + PreservesLimitsOfSize.{w, w'} ((whiskeringLeft C E D).obj F) := + ⟨fun {J} _ => whiskeringLeftPreservesLimitsOfShape J F⟩ instance whiskeringRightPreservesLimitsOfShape {C : Type*} [Category C] {D : Type*} [Category D] {E : Type*} [Category E] {J : Type*} [Category J] From b3184a91682f8ec708b30c9058920bcb8d319c9d Mon Sep 17 00:00:00 2001 From: Robin Carlier Date: Tue, 22 Oct 2024 19:00:57 +0000 Subject: [PATCH 299/505] feat(CategoryTheory/ChosenFiniteProducts): `prodComparison` and `terminalComparison` for `ChosenFiniteProducts` (#17701) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Given a functor `F : C ⥤ D` between categories with chosen finite products and `A B : C`, introduce `terminalComparison F : F (𝟙_ C) ⟶ 𝟙_ D` and `prodComparison F A B : F (A ⊗ B) ⟶ F A ⊗ F B` the canonical maps. - Show that `terminalComparison` is an isomorphism if and only if `F` preserves terminal objects. - Show that the canonical map `prodComparison F A B` is natural in `A` and `B` and compatible with composition. - Show that `F` preserves binary products if and only if `prodComparison F A B` is an isomorphism for all `A B`. Co-authored-by: Robin Carlier <57142648+robin-carlier@users.noreply.github.com> --- .../CategoryTheory/ChosenFiniteProducts.lean | 258 +++++++++++++++++- .../Limits/Shapes/BinaryProducts.lean | 12 + .../Limits/Shapes/Products.lean | 6 + 3 files changed, 274 insertions(+), 2 deletions(-) diff --git a/Mathlib/CategoryTheory/ChosenFiniteProducts.lean b/Mathlib/CategoryTheory/ChosenFiniteProducts.lean index 03e72e08d751c..4fda2a453f164 100644 --- a/Mathlib/CategoryTheory/ChosenFiniteProducts.lean +++ b/Mathlib/CategoryTheory/ChosenFiniteProducts.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2024 Adam Topaz. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Adam Topaz +Authors: Adam Topaz, Robin Carlier -/ import Mathlib.CategoryTheory.Monoidal.OfChosenFiniteProducts.Symmetric import Mathlib.CategoryTheory.Limits.Constructions.FiniteProductsOfBinaryProducts @@ -28,7 +28,7 @@ binary product and `𝟙_ C` for the explicit terminal object. namespace CategoryTheory -universe v u +universe v v₁ v₂ u u₁ u₂ /-- An instance of `ChosenFiniteProducts C` bundles an explicit choice of a binary @@ -228,6 +228,260 @@ instance (priority := 100) : Limits.HasFiniteProducts C := letI : Limits.HasTerminal C := Limits.hasTerminal_of_unique (𝟙_ C) hasFiniteProducts_of_has_binary_and_terminal +section ChosenFiniteProductsComparison + +variable {D : Type u₁} [Category.{v₁} D] [ChosenFiniteProducts D] (F : C ⥤ D) + +section terminalComparison + +/-- When `C` and `D` have chosen finite products and `F : C ⥤ D` is any functor, +`terminalComparison F` is the unique map `F (𝟙_ C) ⟶ 𝟙_ D`. -/ +abbrev terminalComparison : F.obj (𝟙_ C) ⟶ 𝟙_ D := toUnit _ + +@[reassoc (attr := simp)] +lemma map_toUnit_comp_terminalCompariso (A : C) : + F.map (toUnit A) ≫ terminalComparison F = toUnit _ := toUnit_unique _ _ + +open Limits + +/-- If `terminalComparison F` is an Iso, then `F` preserves terminal objects. -/ +noncomputable def preservesLimitEmptyOfIsIsoTerminalComparison [IsIso (terminalComparison F)] : + PreservesLimit (Functor.empty.{0} C) F := by + apply preservesLimitOfPreservesLimitCone terminal.isLimit + apply isLimitChangeEmptyCone D terminal.isLimit + exact asIso (terminalComparison F)|>.symm + +/-- If `F` preserves terminal objects, then `terminalComparison F` is an isomorphism. -/ +def preservesTerminalIso [h : PreservesLimit (Functor.empty.{0} C) F] : F.obj (𝟙_ C) ≅ 𝟙_ D := + (isLimitChangeEmptyCone D (h.preserves terminal.isLimit) (asEmptyCone (F.obj (𝟙_ C))) + (Iso.refl _)).conePointUniqueUpToIso terminal.isLimit + +@[simp] +lemma preservesTerminalIso_hom [PreservesLimit (Functor.empty.{0} C) F] : + (preservesTerminalIso F).hom = terminalComparison F := toUnit_unique _ _ + +instance terminalComparison_isIso_of_preservesLimits [PreservesLimit (Functor.empty.{0} C) F] : + IsIso (terminalComparison F) := by + rw [← preservesTerminalIso_hom] + infer_instance + +end terminalComparison + +section prodComparison + +variable (A B : C) + +/-- When `C` and `D` have chosen finite products and `F : C ⥤ D` is any functor, +`prodComparison F A B` is the canonical comparison morphism from `F (A ⊗ B)` to `F(A) ⊗ F(B)`. -/ +def prodComparison (A B : C) : F.obj (A ⊗ B) ⟶ F.obj A ⊗ F.obj B := + lift (F.map (fst A B)) (F.map (snd A B)) + +@[reassoc (attr := simp)] +theorem prodComparison_fst : prodComparison F A B ≫ fst _ _ = F.map (fst A B) := + lift_fst _ _ + +@[reassoc (attr := simp)] +theorem prodComparison_snd : prodComparison F A B ≫ snd _ _ = F.map (snd A B) := + lift_snd _ _ + +@[reassoc (attr := simp)] +theorem inv_prodComparison_map_fst [IsIso (prodComparison F A B)] : + inv (prodComparison F A B) ≫ F.map (fst _ _) = fst _ _ := by simp [IsIso.inv_comp_eq] + +@[reassoc (attr := simp)] +theorem inv_prodComparison_map_snd [IsIso (prodComparison F A B)] : + inv (prodComparison F A B) ≫ F.map (snd _ _) = snd _ _ := by simp [IsIso.inv_comp_eq] + +variable {A B} {A' B' : C} + +/-- Naturality of the `prodComparison` morphism in both arguments. -/ +@[reassoc] +theorem prodComparison_natural (f : A ⟶ A') (g : B ⟶ B') : + F.map (f ⊗ g) ≫ prodComparison F A' B' = + prodComparison F A B ≫ (F.map f ⊗ F.map g) := by + apply hom_ext <;> + simp only [Category.assoc, prodComparison_fst, tensorHom_fst, prodComparison_fst_assoc, + prodComparison_snd, tensorHom_snd, prodComparison_snd_assoc, ← F.map_comp] + +/-- Naturality of the `prodComparison` morphism in the right argument. -/ +@[reassoc] +theorem prodComparison_natural_whiskerLeft (g : B ⟶ B') : + F.map (A ◁ g) ≫ prodComparison F A B' = + prodComparison F A B ≫ (F.obj A ◁ F.map g) := by + rw [← id_tensorHom, prodComparison_natural, Functor.map_id] + rfl + +/-- Naturality of the `prodComparison` morphism in the left argument. -/ +@[reassoc] +theorem prodComparison_natural_whiskerRight (f : A ⟶ A') : + F.map (f ▷ B) ≫ prodComparison F A' B = + prodComparison F A B ≫ (F.map f ▷ F.obj B) := by + rw [← tensorHom_id, prodComparison_natural, Functor.map_id] + rfl + +section +variable [IsIso (prodComparison F A B)] + +/-- If the product comparison morphism is an iso, its inverse is natural in both argument. -/ +@[reassoc] +theorem prodComparison_inv_natural (f : A ⟶ A') (g : B ⟶ B') [IsIso (prodComparison F A' B')] : + inv (prodComparison F A B) ≫ F.map (f ⊗ g) = + (F.map f ⊗ F.map g) ≫ inv (prodComparison F A' B') := by + rw [IsIso.eq_comp_inv, Category.assoc, IsIso.inv_comp_eq, prodComparison_natural] + +/-- If the product comparison morphism is an iso, its inverse is natural in the right argument. -/ +@[reassoc] +theorem prodComparison_inv_natural_whiskerLeft (g : B ⟶ B') [IsIso (prodComparison F A B')] : + inv (prodComparison F A B) ≫ F.map (A ◁ g) = + (F.obj A ◁ F.map g) ≫ inv (prodComparison F A B') := by + rw [IsIso.eq_comp_inv, Category.assoc, IsIso.inv_comp_eq, prodComparison_natural_whiskerLeft] + +/-- If the product comparison morphism is an iso, its inverse is natural in the left argument. -/ +@[reassoc] +theorem prodComparison_inv_natural_whiskerRight (f : A ⟶ A') [IsIso (prodComparison F A' B)] : + inv (prodComparison F A B) ≫ F.map (f ▷ B) = + (F.map f ▷ F.obj B) ≫ inv (prodComparison F A' B) := by + rw [IsIso.eq_comp_inv, Category.assoc, IsIso.inv_comp_eq, prodComparison_natural_whiskerRight] + +end + +theorem prodComparison_comp {E : Type u₂} [Category.{v₂} E] [ChosenFiniteProducts E] (G : D ⥤ E) : + prodComparison (F ⋙ G) A B = + G.map (prodComparison F A B) ≫ prodComparison G (F.obj A) (F.obj B) := by + unfold prodComparison + ext <;> simp [← G.map_comp] + +@[simp] +lemma prodComparison_id : + prodComparison (𝟭 C) A B = 𝟙 (A ⊗ B) := lift_fst_snd + +/-- The product comparison morphism from `F(A ⊗ -)` to `FA ⊗ F-`, whose components are given by +`prodComparison`. -/ +@[simps] +def prodComparisonNatTrans (A : C) : + (curriedTensor C).obj A ⋙ F ⟶ F ⋙ (curriedTensor D).obj (F.obj A) where + app B := prodComparison F A B + naturality x y f := by + apply hom_ext <;> + simp only [Functor.comp_obj, curriedTensor_obj_obj, + Functor.comp_map, curriedTensor_obj_map, Category.assoc, prodComparison_fst, whiskerLeft_fst, + prodComparison_snd, prodComparison_snd_assoc, whiskerLeft_snd, ← F.map_comp] + +theorem prodComparisonNatTrans_comp {E : Type u₂} [Category.{v₂} E] [ChosenFiniteProducts E] + (G : D ⥤ E) : prodComparisonNatTrans (F ⋙ G) A = + whiskerRight (prodComparisonNatTrans F A) G ≫ + whiskerLeft F (prodComparisonNatTrans G (F.obj A)) := by ext; simp [prodComparison_comp] + +@[simp] +lemma prodComparisonNatTrans_id : + prodComparisonNatTrans (𝟭 C) A = 𝟙 _ := by ext; simp + +/-- The product comparison morphism from `F(- ⊗ -)` to `F- ⊗ F-`, whose components are given by +`prodComparison`. -/ +@[simps] +def prodComparisonBifunctorNatTrans : + curriedTensor C ⋙ (whiskeringRight _ _ _).obj F ⟶ + F ⋙ curriedTensor D ⋙ (whiskeringLeft _ _ _).obj F where + app A := prodComparisonNatTrans F A + naturality x y f := by + ext z + apply hom_ext <;> simp [← Functor.map_comp] + +variable {E : Type u₂} [Category.{v₂} E] + [ChosenFiniteProducts E] (G : D ⥤ E) + +theorem prodComparisonBifunctorNatTrans_comp {E : Type u₂} [Category.{v₂} E] + [ChosenFiniteProducts E] (G : D ⥤ E) : prodComparisonBifunctorNatTrans (F ⋙ G) = + whiskerRight (prodComparisonBifunctorNatTrans F) ((whiskeringRight _ _ _).obj G) ≫ + whiskerLeft F (whiskerRight (prodComparisonBifunctorNatTrans G) + ((whiskeringLeft _ _ _).obj F)) := by ext; simp [prodComparison_comp] + +instance (A : C) [∀ B, IsIso (prodComparison F A B)] : IsIso (prodComparisonNatTrans F A) := by + letI : ∀ X, IsIso ((prodComparisonNatTrans F A).app X) := by assumption + apply NatIso.isIso_of_isIso_app + +instance [∀ A B, IsIso (prodComparison F A B)] : IsIso (prodComparisonBifunctorNatTrans F) := by + letI : ∀ X, IsIso ((prodComparisonBifunctorNatTrans F).app X) := + fun _ ↦ by dsimp; apply NatIso.isIso_of_isIso_app + apply NatIso.isIso_of_isIso_app + +open Limits +section PreservesLimitPairs + +section +variable (A B) +variable [PreservesLimit (pair A B) F] + +/-- If `F` preserves the limit of the pair `(A, B)`, then the binary fan given by +`(F.map fst A B, F.map (snd A B))` is a limit cone. -/ +def isLimitChosenFiniteProductsOfPreservesLimits : + IsLimit <| BinaryFan.mk (F.map (fst A B)) (F.map (snd A B)) := + mapIsLimitOfPreservesOfIsLimit F (fst _ _) (snd _ _) <| + (product A B).isLimit.ofIsoLimit <| isoBinaryFanMk (product A B).cone + +/-- If `F` preserves the limit of the pair `(A, B)`, then `prodComparison F A B` is an isomorphism. +-/ +def prodComparisonIso : F.obj (A ⊗ B) ≅ F.obj A ⊗ F.obj B := + IsLimit.conePointUniqueUpToIso (isLimitChosenFiniteProductsOfPreservesLimits F A B) + (product _ _).isLimit + +@[simp] +lemma prodComparisonIso_hom : (prodComparisonIso F A B).hom = prodComparison F A B := by + rfl + +instance isIso_prodComparison_of_preservesLimit_pair : IsIso (prodComparison F A B) := by + rw [← prodComparisonIso_hom] + infer_instance + +end + +/-- The natural isomorphism `F(A ⊗ -) ≅ FA ⊗ F-`, provided each `prodComparison F A B` is an +isomorphism (as `B` changes). -/ +@[simps! hom inv] +noncomputable def prodComparisonNatIso (A : C) [∀ B, PreservesLimit (pair A B) F] : + (curriedTensor C).obj A ⋙ F ≅ F ⋙ (curriedTensor D).obj (F.obj A) := + asIso (prodComparisonNatTrans F A) + +/-- The natural isomorphism of bifunctors `F(- ⊗ -) ≅ F- ⊗ F-`, provided each +`prodComparison F A B` is an isomorphism. -/ +@[simps! hom inv] +noncomputable def prodComparisonBifunctorNatIso [∀ A B, PreservesLimit (pair A B) F] : + curriedTensor C ⋙ (whiskeringRight _ _ _).obj F ≅ + F ⋙ curriedTensor D ⋙ (whiskeringLeft _ _ _).obj F := + asIso (prodComparisonBifunctorNatTrans F) + +end PreservesLimitPairs + +section ProdComparisonIso + +/-- If `prodComparison F A B` is an isomorphism, then `F` preserves the limit of `pair A B`. -/ +noncomputable def preservesLimitPairOfIsIsoProdComparison (A B : C) [IsIso (prodComparison F A B)] : + PreservesLimit (pair A B) F := by + apply preservesLimitOfPreservesLimitCone (product A B).isLimit + refine IsLimit.equivOfNatIsoOfIso (pairComp A B F) _ + ((product (F.obj A) (F.obj B)).cone.extend (prodComparison F A B)) + (BinaryFan.ext (by exact Iso.refl _) ?_ ?_) |>.invFun + (IsLimit.extendIso _ (product (F.obj A) (F.obj B)).isLimit) + · dsimp only [BinaryFan.fst] + simp [pairComp, prodComparison, lift, fst] + · dsimp only [BinaryFan.snd] + simp [pairComp, prodComparison, lift, snd] + + /-- If `prodComparison F A B` is an isomorphism for all `A B` then `F` preserves limits of shape +`Discrete (WalkingPair)`. -/ +noncomputable def preservesLimitsOfShapeDiscreteWalkingPairOfIsoProdComparison + [∀ A B, IsIso (prodComparison F A B)] : PreservesLimitsOfShape (Discrete WalkingPair) F := by + constructor + intro K + refine @preservesLimitOfIsoDiagram _ _ _ _ _ _ _ _ _ (diagramIsoPair K).symm ?_ + apply preservesLimitPairOfIsIsoProdComparison + +end ProdComparisonIso + +end prodComparison + +end ChosenFiniteProductsComparison + end ChosenFiniteProducts end CategoryTheory diff --git a/Mathlib/CategoryTheory/Limits/Shapes/BinaryProducts.lean b/Mathlib/CategoryTheory/Limits/Shapes/BinaryProducts.lean index c399e2fe79062..841fb4ae1579b 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/BinaryProducts.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/BinaryProducts.lean @@ -196,6 +196,12 @@ theorem BinaryFan.π_app_left {X Y : C} (s : BinaryFan X Y) : s.π.app ⟨Walkin theorem BinaryFan.π_app_right {X Y : C} (s : BinaryFan X Y) : s.π.app ⟨WalkingPair.right⟩ = s.snd := rfl +/-- Constructs an isomorphism of `BinaryFan`s out of an isomorphism of the tips that commutes with +the projections. -/ +def BinaryFan.ext {A B : C} {c c' : BinaryFan A B} (e : c.pt ≅ c'.pt) + (h₁ : c.fst = e.hom ≫ c'.fst) (h₂ : c.snd = e.hom ≫ c'.snd) : c ≅ c' := + Cones.ext e (fun j => by rcases j with ⟨⟨⟩⟩ <;> assumption) + /-- A convenient way to show that a binary fan is a limit. -/ def BinaryFan.IsLimit.mk {X Y : C} (s : BinaryFan X Y) (lift : ∀ {T : C} (_ : T ⟶ X) (_ : T ⟶ Y), T ⟶ s.pt) @@ -225,6 +231,12 @@ abbrev BinaryCofan.inl {X Y : C} (s : BinaryCofan X Y) := s.ι.app ⟨WalkingPai /-- The second inclusion of a binary cofan. -/ abbrev BinaryCofan.inr {X Y : C} (s : BinaryCofan X Y) := s.ι.app ⟨WalkingPair.right⟩ +/-- Constructs an isomorphism of `BinaryCofan`s out of an isomorphism of the tips that commutes with +the injections. -/ +def BinaryCofan.ext {A B : C} {c c' : BinaryCofan A B} (e : c.pt ≅ c'.pt) + (h₁ : c.inl ≫ e.hom = c'.inl) (h₂ : c.inr ≫ e.hom = c'.inr) : c ≅ c' := + Cocones.ext e (fun j => by rcases j with ⟨⟨⟩⟩ <;> assumption) + @[simp] theorem BinaryCofan.ι_app_left {X Y : C} (s : BinaryCofan X Y) : s.ι.app ⟨WalkingPair.left⟩ = s.inl := rfl diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Products.lean b/Mathlib/CategoryTheory/Limits/Shapes/Products.lean index e5ec9c17d3c2a..fd6ac675d5b40 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Products.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Products.lean @@ -245,6 +245,12 @@ theorem Pi.lift_π {β : Type w} {f : β → C} [HasProduct f] {P : C} (p : ∀ Pi.lift p ≫ Pi.π f b = p b := by simp only [limit.lift_π, Fan.mk_pt, Fan.mk_π_app] +/-- A version of `Cones.ext` for `Fan`s. -/ +@[simps!] +def Fan.ext {f : β → C} {c₁ c₂ : Fan f} (e : c₁.pt ≅ c₂.pt) + (w : ∀ (b : β), c₁.proj b = e.hom ≫ c₂.proj b := by aesop_cat) : c₁ ≅ c₂ := + Cones.ext e (fun ⟨j⟩ => w j) + /-- A collection of morphisms `f b ⟶ P` induces a morphism `∐ f ⟶ P`. -/ abbrev Sigma.desc {f : β → C} [HasCoproduct f] {P : C} (p : ∀ b, f b ⟶ P) : ∐ f ⟶ P := colimit.desc _ (Cofan.mk P p) From 58469ae0a68158c2e29247d84339d26ca081d007 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Tue, 22 Oct 2024 19:17:07 +0000 Subject: [PATCH 300/505] chore(Data): delete "`simp` can prove this" porting notes (#18075) This PR checks all porting notes of the form "`simp` can prove this" and if this is indeed the case (and the proof appears to make sense), removes the associated porting note. --- Mathlib/Data/Complex/Abs.lean | 7 ------- Mathlib/Data/Complex/Cardinality.lean | 1 - Mathlib/Data/DFinsupp/Basic.lean | 2 -- Mathlib/Data/Fin/Basic.lean | 3 --- Mathlib/Data/Fin/VecNotation.lean | 3 +-- Mathlib/Data/Finset/Basic.lean | 13 ------------- Mathlib/Data/Finset/Piecewise.lean | 1 - Mathlib/Data/Finsupp/Defs.lean | 1 - Mathlib/Data/Fintype/Card.lean | 2 -- Mathlib/Data/Int/Cast/Lemmas.lean | 10 ---------- Mathlib/Data/Int/Lemmas.lean | 1 - Mathlib/Data/List/Basic.lean | 2 -- Mathlib/Data/List/Cycle.lean | 2 -- Mathlib/Data/List/Intervals.lean | 2 -- Mathlib/Data/Matrix/Basic.lean | 3 --- Mathlib/Data/Matrix/Kronecker.lean | 5 ----- Mathlib/Data/Matrix/Notation.lean | 4 ---- Mathlib/Data/Multiset/Basic.lean | 1 - Mathlib/Data/Multiset/Range.lean | 1 - Mathlib/Data/Nat/Cast/Commute.lean | 7 ------- Mathlib/Data/Nat/Digits.lean | 1 - Mathlib/Data/Nat/Find.lean | 1 - Mathlib/Data/Num/Lemmas.lean | 8 ++++---- Mathlib/Data/Set/Basic.lean | 13 ------------- Mathlib/Data/Set/Finite.lean | 3 --- Mathlib/Data/Set/Lattice.lean | 5 ----- Mathlib/Data/Set/Pointwise/Interval.lean | 9 +++------ Mathlib/Data/Sign.lean | 1 - Mathlib/Data/Sym/Basic.lean | 1 - Mathlib/Data/Vector/Basic.lean | 1 - Mathlib/Data/ZMod/Basic.lean | 1 - 31 files changed, 8 insertions(+), 107 deletions(-) diff --git a/Mathlib/Data/Complex/Abs.lean b/Mathlib/Data/Complex/Abs.lean index ce85a60f52f8c..bd36ccce8e8a2 100644 --- a/Mathlib/Data/Complex/Abs.lean +++ b/Mathlib/Data/Complex/Abs.lean @@ -117,20 +117,13 @@ theorem range_abs : range Complex.abs = Ici 0 := theorem abs_conj (z : ℂ) : Complex.abs (conj z) = Complex.abs z := AbsTheory.abs_conj z --- Porting note (#10618): @[simp] can prove it now theorem abs_prod {ι : Type*} (s : Finset ι) (f : ι → ℂ) : Complex.abs (s.prod f) = s.prod fun I => Complex.abs (f I) := map_prod Complex.abs _ _ --- @[simp] -/- Porting note (#11119): `simp` attribute removed as linter reports this can be proved -by `simp only [@map_pow]` -/ theorem abs_pow (z : ℂ) (n : ℕ) : Complex.abs (z ^ n) = Complex.abs z ^ n := map_pow Complex.abs z n --- @[simp] -/- Porting note (#11119): `simp` attribute removed as linter reports this can be proved -by `simp only [@map_zpow₀]` -/ theorem abs_zpow (z : ℂ) (n : ℤ) : Complex.abs (z ^ n) = Complex.abs z ^ n := map_zpow₀ Complex.abs z n diff --git a/Mathlib/Data/Complex/Cardinality.lean b/Mathlib/Data/Complex/Cardinality.lean index 342a61f0f8696..cb34948f1f671 100644 --- a/Mathlib/Data/Complex/Cardinality.lean +++ b/Mathlib/Data/Complex/Cardinality.lean @@ -24,7 +24,6 @@ theorem mk_complex : #ℂ = 𝔠 := by rw [mk_congr Complex.equivRealProd, mk_prod, lift_id, mk_real, continuum_mul_self] /-- The cardinality of the complex numbers, as a set. -/ --- @[simp] -- Porting note (#10618): simp can prove this theorem mk_univ_complex : #(Set.univ : Set ℂ) = 𝔠 := by rw [mk_univ, mk_complex] /-- The complex numbers are not countable. -/ diff --git a/Mathlib/Data/DFinsupp/Basic.lean b/Mathlib/Data/DFinsupp/Basic.lean index 702ce8b5b5fab..df567e316a5ee 100644 --- a/Mathlib/Data/DFinsupp/Basic.lean +++ b/Mathlib/Data/DFinsupp/Basic.lean @@ -554,7 +554,6 @@ theorem single_apply {i i' b} : theorem single_zero (i) : (single i 0 : Π₀ i, β i) = 0 := DFunLike.coe_injective <| Pi.single_zero _ --- @[simp] -- Porting note (#10618): simp can prove this theorem single_eq_same {i b} : (single i b : Π₀ i, β i) i = b := by simp only [single_apply, dite_eq_ite, ite_true] @@ -655,7 +654,6 @@ def erase (i : ι) (x : Π₀ i, β i) : Π₀ i, β i := theorem erase_apply {i j : ι} {f : Π₀ i, β i} : (f.erase i) j = if j = i then 0 else f j := rfl --- @[simp] -- Porting note (#10618): simp can prove this theorem erase_same {i : ι} {f : Π₀ i, β i} : (f.erase i) i = 0 := by simp theorem erase_ne {i i' : ι} {f : Π₀ i, β i} (h : i' ≠ i) : (f.erase i) i' = f i' := by simp [h] diff --git a/Mathlib/Data/Fin/Basic.lean b/Mathlib/Data/Fin/Basic.lean index 5dbafff567583..95384eac7bdcf 100644 --- a/Mathlib/Data/Fin/Basic.lean +++ b/Mathlib/Data/Fin/Basic.lean @@ -212,10 +212,8 @@ theorem val_fin_lt {n : ℕ} {a b : Fin n} : (a : ℕ) < (b : ℕ) ↔ a < b := theorem val_fin_le {n : ℕ} {a b : Fin n} : (a : ℕ) ≤ (b : ℕ) ↔ a ≤ b := Iff.rfl --- @[simp] -- Porting note (#10618): simp can prove this theorem min_val {a : Fin n} : min (a : ℕ) n = a := by simp --- @[simp] -- Porting note (#10618): simp can prove this theorem max_val {a : Fin n} : max (a : ℕ) n = n := by simp /-- The inclusion map `Fin n → ℕ` is an embedding. -/ @@ -408,7 +406,6 @@ theorem nontrivial_iff_two_le : Nontrivial (Fin n) ↔ 2 ≤ n := by section Monoid --- Porting note (#10618): removing `simp`, `simp` can prove it with AddCommMonoid instance protected theorem add_zero [NeZero n] (k : Fin n) : k + 0 = k := by simp only [add_def, val_zero', Nat.add_zero, mod_eq_of_lt (is_lt k)] diff --git a/Mathlib/Data/Fin/VecNotation.lean b/Mathlib/Data/Fin/VecNotation.lean index c931078da5ce1..c61d6aab904ce 100644 --- a/Mathlib/Data/Fin/VecNotation.lean +++ b/Mathlib/Data/Fin/VecNotation.lean @@ -160,11 +160,10 @@ theorem range_cons (x : α) (u : Fin n → α) : Set.range (vecCons x u) = {x} theorem range_empty (u : Fin 0 → α) : Set.range u = ∅ := Set.range_eq_empty _ --- @[simp] -- Porting note (#10618): simp can prove this theorem range_cons_empty (x : α) (u : Fin 0 → α) : Set.range (Matrix.vecCons x u) = {x} := by rw [range_cons, range_empty, Set.union_empty] --- @[simp] -- Porting note (#10618): simp can prove this (up to commutativity) +-- simp can prove this (up to commutativity) theorem range_cons_cons_empty (x y : α) (u : Fin 0 → α) : Set.range (vecCons x <| vecCons y u) = {x, y} := by rw [range_cons, range_cons_empty, Set.singleton_union] diff --git a/Mathlib/Data/Finset/Basic.lean b/Mathlib/Data/Finset/Basic.lean index a94497b850efa..8dad68581aed6 100644 --- a/Mathlib/Data/Finset/Basic.lean +++ b/Mathlib/Data/Finset/Basic.lean @@ -206,7 +206,6 @@ theorem setOf_mem {α} {s : Finset α} : { a | a ∈ s } = s := theorem coe_mem {s : Finset α} (x : (s : Set α)) : ↑x ∈ s := x.2 --- Porting note (#10618): @[simp] can prove this theorem mk_coe {s : Finset α} (x : (s : Set α)) {h} : (⟨x, h⟩ : (s : Set α)) = x := Subtype.coe_eta _ _ @@ -232,12 +231,10 @@ theorem coe_injective {α} : Injective ((↑) : Finset α → Set α) := fun _s instance {α : Type u} : CoeSort (Finset α) (Type u) := ⟨fun s => { x // x ∈ s }⟩ --- Porting note (#10618): @[simp] can prove this protected theorem forall_coe {α : Type*} (s : Finset α) (p : s → Prop) : (∀ x : s, p x) ↔ ∀ (x : α) (h : x ∈ s), p ⟨x, h⟩ := Subtype.forall --- Porting note (#10618): @[simp] can prove this protected theorem exists_coe {α : Type*} (s : Finset α) (p : s → Prop) : (∃ x : s, p x) ↔ ∃ (x : α) (h : x ∈ s), p ⟨x, h⟩ := Subtype.exists @@ -747,7 +744,6 @@ theorem mem_cons {h} : b ∈ s.cons a h ↔ b = a ∨ b ∈ s := theorem mem_cons_of_mem {a b : α} {s : Finset α} {hb : b ∉ s} (ha : a ∈ s) : a ∈ cons b s hb := Multiset.mem_cons_of_mem ha --- Porting note (#10618): @[simp] can prove this theorem mem_cons_self (a : α) (s : Finset α) {h} : a ∈ cons a s h := Multiset.mem_cons_self _ _ @@ -991,14 +987,12 @@ theorem insert_eq_self : insert a s = s ↔ a ∈ s := theorem insert_ne_self : insert a s ≠ s ↔ a ∉ s := insert_eq_self.not --- Porting note (#10618): @[simp] can prove this theorem pair_eq_singleton (a : α) : ({a, a} : Finset α) = {a} := insert_eq_of_mem <| mem_singleton_self _ theorem Insert.comm (a b : α) (s : Finset α) : insert a (insert b s) = insert b (insert a s) := ext fun x => by simp only [mem_insert, or_left_comm] --- Porting note (#10618): @[simp] can prove this @[norm_cast] theorem coe_pair {a b : α} : (({a, b} : Finset α) : Set α) = {a, b} := by ext @@ -1011,7 +1005,6 @@ theorem coe_eq_pair {s : Finset α} {a b : α} : (s : Set α) = {a, b} ↔ s = { theorem pair_comm (a b : α) : ({a, b} : Finset α) = {b, a} := Insert.comm a b ∅ --- Porting note (#10618): @[simp] can prove this theorem insert_idem (a : α) (s : Finset α) : insert a (insert a s) = insert a s := ext fun x => by simp only [mem_insert, ← or_assoc, or_self_iff] @@ -1493,13 +1486,11 @@ instance : DistribLattice (Finset α) := @[simp] theorem union_left_idem (s t : Finset α) : s ∪ (s ∪ t) = s ∪ t := sup_left_idem _ _ --- Porting note (#10618): @[simp] can prove this theorem union_right_idem (s t : Finset α) : s ∪ t ∪ t = s ∪ t := sup_right_idem _ _ @[simp] theorem inter_left_idem (s t : Finset α) : s ∩ (s ∩ t) = s ∩ t := inf_left_idem _ _ --- Porting note (#10618): @[simp] can prove this theorem inter_right_idem (s t : Finset α) : s ∩ t ∩ t = s ∩ t := inf_right_idem _ _ theorem inter_union_distrib_left (s t u : Finset α) : s ∩ (t ∪ u) = s ∩ t ∪ s ∩ u := @@ -1805,7 +1796,6 @@ theorem inter_sdiff (s t u : Finset α) : s ∩ (t \ u) = (s ∩ t) \ u := (inte theorem sdiff_inter_self (s₁ s₂ : Finset α) : s₂ \ s₁ ∩ s₁ = ∅ := inf_sdiff_self_left --- Porting note (#10618): @[simp] can prove this protected theorem sdiff_self (s₁ : Finset α) : s₁ \ s₁ = ∅ := _root_.sdiff_self @@ -1861,7 +1851,6 @@ theorem union_sdiff_symm : s ∪ t \ s = t ∪ s \ t := by simp [union_comm] theorem sdiff_union_inter (s t : Finset α) : s \ t ∪ s ∩ t = s := sup_sdiff_inf _ _ --- Porting note (#10618): @[simp] can prove this theorem sdiff_idem (s t : Finset α) : (s \ t) \ t = s \ t := _root_.sdiff_idem @@ -2553,11 +2542,9 @@ theorem range_succ : range (succ n) = insert n (range n) := theorem range_add_one : range (n + 1) = insert n (range n) := range_succ --- Porting note (#10618): @[simp] can prove this theorem not_mem_range_self : n ∉ range n := Multiset.not_mem_range_self --- Porting note (#10618): @[simp] can prove this theorem self_mem_range_succ (n : ℕ) : n ∈ range (n + 1) := Multiset.self_mem_range_succ n diff --git a/Mathlib/Data/Finset/Piecewise.lean b/Mathlib/Data/Finset/Piecewise.lean index 959a4e4280a6e..99cea5cc557d0 100644 --- a/Mathlib/Data/Finset/Piecewise.lean +++ b/Mathlib/Data/Finset/Piecewise.lean @@ -26,7 +26,6 @@ variable {ι : Type*} {π : ι → Sort*} (s : Finset ι) (f g : ∀ i, π i) complement. -/ def piecewise [∀ j, Decidable (j ∈ s)] : ∀ i, π i := fun i ↦ if i ∈ s then f i else g i --- Porting note (#10618): @[simp] can prove this lemma piecewise_insert_self [DecidableEq ι] {j : ι} [∀ i, Decidable (i ∈ insert j s)] : (insert j s).piecewise f g j = f j := by simp [piecewise] diff --git a/Mathlib/Data/Finsupp/Defs.lean b/Mathlib/Data/Finsupp/Defs.lean index c424ca7128840..b44c8ceb8c1bb 100644 --- a/Mathlib/Data/Finsupp/Defs.lean +++ b/Mathlib/Data/Finsupp/Defs.lean @@ -637,7 +637,6 @@ theorem support_onFinset_subset {s : Finset α} {f : α → M} {hf} : (onFinset s f hf).support ⊆ s := by classical convert filter_subset (f · ≠ 0) s --- @[simp] -- Porting note (#10618): simp can prove this theorem mem_support_onFinset {s : Finset α} {f : α → M} (hf : ∀ a : α, f a ≠ 0 → a ∈ s) {a : α} : a ∈ (Finsupp.onFinset s f hf).support ↔ f a ≠ 0 := by rw [Finsupp.mem_support_iff, Finsupp.onFinset_apply] diff --git a/Mathlib/Data/Fintype/Card.lean b/Mathlib/Data/Fintype/Card.lean index 641354c5ca1ae..48349c1a8a9b7 100644 --- a/Mathlib/Data/Fintype/Card.lean +++ b/Mathlib/Data/Fintype/Card.lean @@ -306,12 +306,10 @@ theorem Fin.cast_eq_cast' {n m : ℕ} (h : Fin n = Fin m) : theorem card_finset_fin_le {n : ℕ} (s : Finset (Fin n)) : s.card ≤ n := by simpa only [Fintype.card_fin] using s.card_le_univ ---@[simp] Porting note (#10618): simp can prove it theorem Fintype.card_subtype_eq (y : α) [Fintype { x // x = y }] : Fintype.card { x // x = y } = 1 := Fintype.card_unique ---@[simp] Porting note (#10618): simp can prove it theorem Fintype.card_subtype_eq' (y : α) [Fintype { x // y = x }] : Fintype.card { x // y = x } = 1 := Fintype.card_unique diff --git a/Mathlib/Data/Int/Cast/Lemmas.lean b/Mathlib/Data/Int/Cast/Lemmas.lean index 822449e1e3d3b..ef53e50244441 100644 --- a/Mathlib/Data/Int/Cast/Lemmas.lean +++ b/Mathlib/Data/Int/Cast/Lemmas.lean @@ -172,18 +172,8 @@ lemma intCast_mul_intCast_mul (h : Commute a b) (m n : ℤ) : Commute (m * a) (n variable (a) (m n : ℤ) -/- Porting note (#10618): `simp` attribute removed as linter reports: -simp can prove this: - by simp only [Commute.cast_int_right, Commute.refl, Commute.mul_right] --/ --- @[simp] lemma self_intCast_mul : Commute a (n * a : α) := (Commute.refl a).intCast_mul_right n -/- Porting note (#10618): `simp` attribute removed as linter reports: -simp can prove this: - by simp only [Commute.cast_int_left, Commute.refl, Commute.mul_left] --/ --- @[simp] lemma intCast_mul_self : Commute ((n : α) * a) a := (Commute.refl a).intCast_mul_left n lemma self_intCast_mul_intCast_mul : Commute (m * a : α) (n * a : α) := diff --git a/Mathlib/Data/Int/Lemmas.lean b/Mathlib/Data/Int/Lemmas.lean index 3842fbb1c3222..7808804746ccb 100644 --- a/Mathlib/Data/Int/Lemmas.lean +++ b/Mathlib/Data/Int/Lemmas.lean @@ -30,7 +30,6 @@ theorem le_natCast_sub (m n : ℕ) : (m - n : ℤ) ≤ ↑(m - n : ℕ) := by /-! ### `succ` and `pred` -/ --- Porting note (#10618): simp can prove this @[simp] theorem succ_natCast_pos (n : ℕ) : 0 < (n : ℤ) + 1 := lt_add_one_iff.mpr (by simp) diff --git a/Mathlib/Data/List/Basic.lean b/Mathlib/Data/List/Basic.lean index 3f6d1c26e9f52..74c813e67ec38 100644 --- a/Mathlib/Data/List/Basic.lean +++ b/Mathlib/Data/List/Basic.lean @@ -1113,8 +1113,6 @@ theorem foldl_fixed {a : α} : ∀ l : List β, foldl (fun a _ => a) a l = a := theorem foldr_fixed {b : β} : ∀ l : List α, foldr (fun _ b => b) b l = b := foldr_fixed' fun _ => rfl --- Porting note (#10618): simp can prove this --- @[simp] theorem foldr_eta : ∀ l : List α, foldr cons [] l = l := by simp only [foldr_cons_eq_append, append_nil, forall_const] diff --git a/Mathlib/Data/List/Cycle.lean b/Mathlib/Data/List/Cycle.lean index 18607fde71b44..2f3c103975e43 100644 --- a/Mathlib/Data/List/Cycle.lean +++ b/Mathlib/Data/List/Cycle.lean @@ -196,7 +196,6 @@ theorem prev_getLast_cons (h : x ∈ x :: l) : theorem prev_cons_cons_eq' (y z : α) (h : x ∈ y :: z :: l) (hx : x = y) : prev (y :: z :: l) x h = getLast (z :: l) (cons_ne_nil _ _) := by rw [prev, dif_pos hx] ---@[simp] Porting note (#10618): `simp` can prove it theorem prev_cons_cons_eq (z : α) (h : x ∈ x :: z :: l) : prev (x :: z :: l) x h = getLast (z :: l) (cons_ne_nil _ _) := prev_cons_cons_eq' l x x z h rfl @@ -812,7 +811,6 @@ theorem chain_coe_cons (r : α → α → Prop) (a : α) (l : List α) : Chain r (a :: l) ↔ List.Chain r a (l ++ [a]) := Iff.rfl ---@[simp] Porting note (#10618): `simp` can prove it theorem chain_singleton (r : α → α → Prop) (a : α) : Chain r [a] ↔ r a a := by rw [chain_coe_cons, nil_append, List.chain_singleton] diff --git a/Mathlib/Data/List/Intervals.lean b/Mathlib/Data/List/Intervals.lean index 870f434107df2..9c89237d400af 100644 --- a/Mathlib/Data/List/Intervals.lean +++ b/Mathlib/Data/List/Intervals.lean @@ -125,8 +125,6 @@ theorem chain'_succ (n m : ℕ) : Chain' (fun a b => b = succ a) (Ico n m) := by · rw [eq_nil_of_le (le_of_not_gt h)] trivial --- Porting note (#10618): simp can prove this --- @[simp] theorem not_mem_top {n m : ℕ} : m ∉ Ico n m := by simp theorem filter_lt_of_top_le {n m l : ℕ} (hml : m ≤ l) : diff --git a/Mathlib/Data/Matrix/Basic.lean b/Mathlib/Data/Matrix/Basic.lean index 551793cf8e46b..6194d386d0f50 100644 --- a/Mathlib/Data/Matrix/Basic.lean +++ b/Mathlib/Data/Matrix/Basic.lean @@ -2111,7 +2111,6 @@ theorem conjTranspose_smul_non_comm [Star R] [Star α] [SMul R α] [SMul Rᵐᵒ (c • M)ᴴ = MulOpposite.op (star c) • Mᴴ := Matrix.ext <| by simp [h] --- @[simp] -- Porting note (#10618): simp can prove this theorem conjTranspose_smul_self [Mul α] [StarMul α] (c : α) (M : Matrix m n α) : (c • M)ᴴ = MulOpposite.op (star c) • Mᴴ := conjTranspose_smul_non_comm c M star_mul @@ -2449,7 +2448,6 @@ theorem reindex_apply (eₘ : m ≃ l) (eₙ : n ≃ o) (M : Matrix m n α) : reindex eₘ eₙ M = M.submatrix eₘ.symm eₙ.symm := rfl --- @[simp] -- Porting note (#10618): simp can prove this theorem reindex_refl_refl (A : Matrix m n α) : reindex (Equiv.refl _) (Equiv.refl _) A = A := A.submatrix_id_id @@ -2472,7 +2470,6 @@ theorem conjTranspose_reindex [Star α] (eₘ : m ≃ l) (eₙ : n ≃ o) (M : M (reindex eₘ eₙ M)ᴴ = reindex eₙ eₘ Mᴴ := rfl --- @[simp] -- Porting note (#10618): simp can prove this theorem submatrix_mul_transpose_submatrix [Fintype m] [Fintype n] [AddCommMonoid α] [Mul α] (e : m ≃ n) (M : Matrix m n α) : M.submatrix id e * Mᵀ.submatrix e id = M * Mᵀ := by rw [submatrix_mul_equiv, submatrix_id_id] diff --git a/Mathlib/Data/Matrix/Kronecker.lean b/Mathlib/Data/Matrix/Kronecker.lean index 61bfd0d4ab97b..9c1172d6755c4 100644 --- a/Mathlib/Data/Matrix/Kronecker.lean +++ b/Mathlib/Data/Matrix/Kronecker.lean @@ -259,11 +259,9 @@ def kroneckerBilinear [CommSemiring R] [Semiring α] [Algebra R α] : hypotheses which can be filled by properties of `*`. -/ --- @[simp] -- Porting note (#10618): simp can prove this theorem zero_kronecker [MulZeroClass α] (B : Matrix n p α) : (0 : Matrix l m α) ⊗ₖ B = 0 := kroneckerMap_zero_left _ zero_mul B --- @[simp] -- Porting note (#10618): simp can prove this theorem kronecker_zero [MulZeroClass α] (A : Matrix l m α) : A ⊗ₖ (0 : Matrix n p α) = 0 := kroneckerMap_zero_right _ mul_zero A @@ -327,7 +325,6 @@ theorem ofNat_kronecker [Semiring α] [DecidableEq l] (a : ℕ) [a.AtLeastTwo] ( (blockDiagonal fun _ => (OfNat.ofNat a : α) • B) := diagonal_kronecker _ _ --- @[simp] -- Porting note (#10618): simp can prove this theorem one_kronecker_one [MulZeroOneClass α] [DecidableEq m] [DecidableEq n] : (1 : Matrix m m α) ⊗ₖ (1 : Matrix n n α) = 1 := kroneckerMap_one_one _ zero_mul mul_zero (one_mul _) @@ -443,11 +440,9 @@ def kroneckerTMulBilinear : hypotheses which can be filled by properties of `⊗ₜ`. -/ --- @[simp] -- Porting note (#10618): simp can prove this theorem zero_kroneckerTMul (B : Matrix n p β) : (0 : Matrix l m α) ⊗ₖₜ[R] B = 0 := kroneckerMap_zero_left _ (zero_tmul α) B --- @[simp] -- Porting note (#10618): simp can prove this theorem kroneckerTMul_zero (A : Matrix l m α) : A ⊗ₖₜ[R] (0 : Matrix n p β) = 0 := kroneckerMap_zero_right _ (tmul_zero β) A diff --git a/Mathlib/Data/Matrix/Notation.lean b/Mathlib/Data/Matrix/Notation.lean index 34515d69bf4ca..82b807cb9664f 100644 --- a/Mathlib/Data/Matrix/Notation.lean +++ b/Mathlib/Data/Matrix/Notation.lean @@ -180,7 +180,6 @@ theorem dotProduct_cons (v : Fin n.succ → α) (x : α) (w : Fin n → α) : dotProduct v (vecCons x w) = vecHead v * x + dotProduct (vecTail v) w := by simp [dotProduct, Fin.sum_univ_succ, vecHead, vecTail] --- @[simp] -- Porting note (#10618): simp can prove this theorem cons_dotProduct_cons (x : α) (v : Fin n → α) (y : α) (w : Fin n → α) : dotProduct (vecCons x v) (vecCons y w) = x * y + dotProduct v w := by simp @@ -292,7 +291,6 @@ theorem vecMul_cons (v : Fin n.succ → α) (w : o' → α) (B : Fin n → o' ext i simp [vecMul] --- @[simp] -- Porting note (#10618): simp can prove this theorem cons_vecMul_cons (x : α) (v : Fin n → α) (w : o' → α) (B : Fin n → o' → α) : vecCons x v ᵥ* of (vecCons w B) = x • w + v ᵥ* of B := by simp @@ -352,11 +350,9 @@ section SMul variable [NonUnitalNonAssocSemiring α] --- @[simp] -- Porting note (#10618): simp can prove this theorem smul_mat_empty {m' : Type*} (x : α) (A : Fin 0 → m' → α) : x • A = ![] := empty_eq _ --- @[simp] -- Porting note (#10618): simp can prove this theorem smul_mat_cons (x : α) (v : n' → α) (A : Fin m → n' → α) : x • vecCons v A = vecCons (x • v) (x • A) := by ext i diff --git a/Mathlib/Data/Multiset/Basic.lean b/Mathlib/Data/Multiset/Basic.lean index 97366784110f6..001ad7a8b6abd 100644 --- a/Mathlib/Data/Multiset/Basic.lean +++ b/Mathlib/Data/Multiset/Basic.lean @@ -226,7 +226,6 @@ theorem mem_cons {a b : α} {s : Multiset α} : a ∈ b ::ₘ s ↔ a = b ∨ a theorem mem_cons_of_mem {a b : α} {s : Multiset α} (h : a ∈ s) : a ∈ b ::ₘ s := mem_cons.2 <| Or.inr h --- @[simp] -- Porting note (#10618): simp can prove this theorem mem_cons_self (a : α) (s : Multiset α) : a ∈ a ::ₘ s := mem_cons.2 (Or.inl rfl) diff --git a/Mathlib/Data/Multiset/Range.lean b/Mathlib/Data/Multiset/Range.lean index f43a5a4f36091..ef904f2a33533 100644 --- a/Mathlib/Data/Multiset/Range.lean +++ b/Mathlib/Data/Multiset/Range.lean @@ -40,7 +40,6 @@ theorem range_subset {m n : ℕ} : range m ⊆ range n ↔ m ≤ n := theorem mem_range {m n : ℕ} : m ∈ range n ↔ m < n := List.mem_range --- Porting note (#10618): removing @[simp], `simp` can prove it theorem not_mem_range_self {n : ℕ} : n ∉ range n := List.not_mem_range_self diff --git a/Mathlib/Data/Nat/Cast/Commute.lean b/Mathlib/Data/Nat/Cast/Commute.lean index 9d1e9bc89094b..a2d80c1d27b64 100644 --- a/Mathlib/Data/Nat/Cast/Commute.lean +++ b/Mathlib/Data/Nat/Cast/Commute.lean @@ -71,17 +71,10 @@ variable [Semiring α] {a b : α} variable (a) (m n : ℕ) --- Porting note (#10618): `simp` can prove this using `Commute.refl`, `Commute.natCast_mul_right` --- @[simp] lemma self_natCast_mul : Commute a (n * a) := (Commute.refl a).natCast_mul_right n --- Porting note (#10618): `simp` can prove this using `Commute.refl`, `Commute.natCast_mul_left` --- @[simp] lemma natCast_mul_self : Commute (n * a) a := (Commute.refl a).natCast_mul_left n --- Porting note (#10618): `simp` can prove this using `Commute.refl`, `Commute.natCast_mul_left`, --- `Commute.natCast_mul_right` --- @[simp] lemma self_natCast_mul_natCast_mul : Commute (m * a) (n * a) := (Commute.refl a).natCast_mul_natCast_mul m n diff --git a/Mathlib/Data/Nat/Digits.lean b/Mathlib/Data/Nat/Digits.lean index 23cd40d657cba..2af99d2651db3 100644 --- a/Mathlib/Data/Nat/Digits.lean +++ b/Mathlib/Data/Nat/Digits.lean @@ -82,7 +82,6 @@ def digits : ℕ → ℕ → List ℕ theorem digits_zero (b : ℕ) : digits b 0 = [] := by rcases b with (_ | ⟨_ | ⟨_⟩⟩) <;> simp [digits, digitsAux0, digitsAux1] --- @[simp] -- Porting note (#10618): simp can prove this theorem digits_zero_zero : digits 0 0 = [] := rfl diff --git a/Mathlib/Data/Nat/Find.lean b/Mathlib/Data/Nat/Find.lean index b26536ee0f517..a86b825b1b788 100644 --- a/Mathlib/Data/Nat/Find.lean +++ b/Mathlib/Data/Nat/Find.lean @@ -106,7 +106,6 @@ lemma find_comp_succ (h₁ : ∃ n, p n) (h₂ : ∃ n, p (n + 1)) (h0 : ¬ p 0) cases n exacts [h0, @Nat.find_min (fun n ↦ p (n + 1)) _ h₂ _ (succ_lt_succ_iff.1 hn)] --- Porting note (#10618): removing `simp` attribute as `simp` can prove it lemma find_pos (h : ∃ n : ℕ, p n) : 0 < Nat.find h ↔ ¬p 0 := Nat.pos_iff_ne_zero.trans (Nat.find_eq_zero _).not diff --git a/Mathlib/Data/Num/Lemmas.lean b/Mathlib/Data/Num/Lemmas.lean index 1ee383be417b7..1d1f99cc29e6a 100644 --- a/Mathlib/Data/Num/Lemmas.lean +++ b/Mathlib/Data/Num/Lemmas.lean @@ -399,11 +399,11 @@ instance linearOrderedSemiring : LinearOrderedSemiring Num := decidableEq := instDecidableEqNum exists_pair_ne := ⟨0, 1, by decide⟩ } -@[norm_cast] -- @[simp] -- Porting note (#10618): simp can prove this +@[norm_cast] theorem add_of_nat (m n) : ((m + n : ℕ) : Num) = m + n := add_ofNat' _ _ -@[norm_cast] -- @[simp] -- Porting note (#10618): simp can prove this +@[norm_cast] theorem to_nat_to_int (n : Num) : ((n : ℕ) : ℤ) = n := cast_to_nat _ @@ -422,7 +422,7 @@ theorem of_natCast {α} [AddMonoidWithOne α] (n : ℕ) : ((n : Num) : α) = n : @[deprecated (since := "2024-04-17")] alias of_nat_cast := of_natCast -@[norm_cast] -- @[simp] -- Porting note (#10618): simp can prove this +@[norm_cast] theorem of_nat_inj {m n : ℕ} : (m : Num) = n ↔ m = n := ⟨fun h => Function.LeftInverse.injective to_of_nat h, congr_arg _⟩ @@ -1323,7 +1323,7 @@ instance linearOrderedCommRing : LinearOrderedCommRing ZNum := @[simp, norm_cast] theorem cast_sub [Ring α] (m n) : ((m - n : ZNum) : α) = m - n := by simp [sub_eq_neg_add] -@[norm_cast] -- @[simp] -- Porting note (#10618): simp can prove this +@[norm_cast] theorem neg_of_int : ∀ n, ((-n : ℤ) : ZNum) = -n | (_ + 1 : ℕ) => rfl | 0 => by rw [Int.cast_neg] diff --git a/Mathlib/Data/Set/Basic.lean b/Mathlib/Data/Set/Basic.lean index 6c5861316434a..f755b5a1cb2a8 100644 --- a/Mathlib/Data/Set/Basic.lean +++ b/Mathlib/Data/Set/Basic.lean @@ -142,11 +142,9 @@ theorem Set.coe_eq_subtype (s : Set α) : ↥s = { x // x ∈ s } := theorem Set.coe_setOf (p : α → Prop) : ↥{ x | p x } = { x // p x } := rfl --- Porting note (#10618): removed `simp` because `simp` can prove it theorem SetCoe.forall {s : Set α} {p : s → Prop} : (∀ x : s, p x) ↔ ∀ (x) (h : x ∈ s), p ⟨x, h⟩ := Subtype.forall --- Porting note (#10618): removed `simp` because `simp` can prove it theorem SetCoe.exists {s : Set α} {p : s → Prop} : (∃ x : s, p x) ↔ ∃ (x : _) (h : x ∈ s), p ⟨x, h⟩ := Subtype.exists @@ -347,7 +345,6 @@ protected theorem ssubset_of_subset_of_ssubset {s₁ s₂ s₃ : Set α} (hs₁s theorem not_mem_empty (x : α) : ¬x ∈ (∅ : Set α) := id --- Porting note (#10618): removed `simp` because `simp` can prove it theorem not_not_mem : ¬a ∉ s ↔ a ∈ s := not_not @@ -355,8 +352,6 @@ theorem not_not_mem : ¬a ∉ s ↔ a ∈ s := -- Porting note: we seem to need parentheses at `(↥s)`, -- even if we increase the right precedence of `↥` in `Mathlib.Tactic.Coe`. --- Porting note: removed `simp` as it is competing with `nonempty_subtype`. --- @[simp] theorem nonempty_coe_sort {s : Set α} : Nonempty (↥s) ↔ s.Nonempty := nonempty_subtype @@ -942,7 +937,6 @@ theorem ssubset_insert {s : Set α} {a : α} (h : a ∉ s) : s ⊂ insert a s := theorem insert_comm (a b : α) (s : Set α) : insert a (insert b s) = insert b (insert a s) := ext fun _ => or_left_comm --- Porting note (#10618): removing `simp` attribute because `simp` can prove it theorem insert_idem (a : α) (s : Set α) : insert a (insert a s) = insert a s := insert_eq_of_mem <| mem_insert _ _ @@ -1046,7 +1040,6 @@ theorem singleton_nonempty (a : α) : ({a} : Set α).Nonempty := theorem singleton_ne_empty (a : α) : ({a} : Set α) ≠ ∅ := (singleton_nonempty _).ne_empty ---Porting note (#10618): removed `simp` attribute because `simp` can prove it theorem empty_ssubset_singleton : (∅ : Set α) ⊂ {a} := (singleton_nonempty _).empty_ssubset @@ -1136,19 +1129,15 @@ theorem sep_eq_self_iff_mem_true : { x ∈ s | p x } = s ↔ ∀ x ∈ s, p x := theorem sep_eq_empty_iff_mem_false : { x ∈ s | p x } = ∅ ↔ ∀ x ∈ s, ¬p x := by simp_rw [Set.ext_iff, mem_sep_iff, mem_empty_iff_false, iff_false, not_and] ---Porting note (#10618): removed `simp` attribute because `simp` can prove it theorem sep_true : { x ∈ s | True } = s := inter_univ s ---Porting note (#10618): removed `simp` attribute because `simp` can prove it theorem sep_false : { x ∈ s | False } = ∅ := inter_empty s ---Porting note (#10618): removed `simp` attribute because `simp` can prove it theorem sep_empty (p : α → Prop) : { x ∈ (∅ : Set α) | p x } = ∅ := empty_inter {x | p x} ---Porting note (#10618): removed `simp` attribute because `simp` can prove it theorem sep_univ : { x ∈ (univ : Set α) | p x } = { x | p x } := univ_inter {x | p x} @@ -1643,7 +1632,6 @@ theorem insert_diff_singleton_comm (hab : a ≠ b) (s : Set α) : simp_rw [← union_singleton, union_diff_distrib, diff_singleton_eq_self (mem_singleton_iff.not.2 hab.symm)] ---Porting note (#10618): removed `simp` attribute because `simp` can prove it theorem diff_self {s : Set α} : s \ s = ∅ := sdiff_self @@ -1671,7 +1659,6 @@ theorem union_eq_diff_union_diff_union_inter (s t : Set α) : s ∪ t = s \ t /-! ### Lemmas about pairs -/ ---Porting note (#10618): removed `simp` attribute because `simp` can prove it theorem pair_eq_singleton (a : α) : ({a, a} : Set α) = {a} := union_self _ diff --git a/Mathlib/Data/Set/Finite.lean b/Mathlib/Data/Set/Finite.lean index cdbea2da57322..b6f134607f7fc 100644 --- a/Mathlib/Data/Set/Finite.lean +++ b/Mathlib/Data/Set/Finite.lean @@ -266,8 +266,6 @@ protected theorem toFinset_image [DecidableEq β] (f : α → β) (hs : s.Finite ext simp --- Porting note (#10618): now `simp` can prove it but it needs the `fintypeRange` instance --- from the next section protected theorem toFinset_range [DecidableEq α] [Fintype β] (f : β → α) (h : (range f).Finite) : h.toFinset = Finset.univ.image f := by ext @@ -482,7 +480,6 @@ This is a wrapper around `Set.toFinite`. -/ theorem finite_toSet (s : Finset α) : (s : Set α).Finite := Set.toFinite _ --- Porting note (#10618): was @[simp], now `simp` can prove it theorem finite_toSet_toFinset (s : Finset α) : s.finite_toSet.toFinset = s := by rw [toFinite_toFinset, toFinset_coe] diff --git a/Mathlib/Data/Set/Lattice.lean b/Mathlib/Data/Set/Lattice.lean index 4892bed5f539c..2682e7dba7ab8 100644 --- a/Mathlib/Data/Set/Lattice.lean +++ b/Mathlib/Data/Set/Lattice.lean @@ -239,7 +239,6 @@ theorem subset_iInter_iff {s : Set α} {t : ι → Set α} : (s ⊆ ⋂ i, t i) le_iInf_iff /- ./././Mathport/Syntax/Translate/Expr.lean:107:6: warning: expanding binder group (i j) -/ --- Porting note (#10618): removing `simp`. `simp` can prove it theorem subset_iInter₂_iff {s : Set α} {t : ∀ i, κ i → Set α} : (s ⊆ ⋂ (i) (j), t i j) ↔ ∀ i j, s ⊆ t i j := by simp_rw [subset_iInter_iff] @@ -568,7 +567,6 @@ theorem iInter_eq_univ : ⋂ i, s i = univ ↔ ∀ i, s i = univ := theorem nonempty_iUnion : (⋃ i, s i).Nonempty ↔ ∃ i, (s i).Nonempty := by simp [nonempty_iff_ne_empty] --- Porting note (#10618): removing `simp`. `simp` can prove it theorem nonempty_biUnion {t : Set α} {s : α → Set β} : (⋃ i ∈ t, s i).Nonempty ↔ ∃ i ∈ t, (s i).Nonempty := by simp @@ -1019,7 +1017,6 @@ theorem nonempty_iInter {f : ι → Set α} : (⋂ i, f i).Nonempty ↔ ∃ x, /- ./././Mathport/Syntax/Translate/Expr.lean:107:6: warning: expanding binder group (i j) -/ -- classical --- Porting note (#10618): removing `simp`. `simp` can prove it theorem nonempty_iInter₂ {s : ∀ i, κ i → Set α} : (⋂ (i) (j), s i j).Nonempty ↔ ∃ a, ∀ i j, a ∈ s i j := by simp @@ -1767,13 +1764,11 @@ theorem disjoint_iUnion_right {ι : Sort*} {s : ι → Set α} : disjoint_iSup_iff /- ./././Mathport/Syntax/Translate/Expr.lean:107:6: warning: expanding binder group (i j) -/ --- Porting note (#10618): removing `simp`. `simp` can prove it theorem disjoint_iUnion₂_left {s : ∀ i, κ i → Set α} {t : Set α} : Disjoint (⋃ (i) (j), s i j) t ↔ ∀ i j, Disjoint (s i j) t := iSup₂_disjoint_iff /- ./././Mathport/Syntax/Translate/Expr.lean:107:6: warning: expanding binder group (i j) -/ --- Porting note (#10618): removing `simp`. `simp` can prove it theorem disjoint_iUnion₂_right {s : Set α} {t : ∀ i, κ i → Set α} : Disjoint s (⋃ (i) (j), t i j) ↔ ∀ i j, Disjoint s (t i j) := disjoint_iSup₂_iff diff --git a/Mathlib/Data/Set/Pointwise/Interval.lean b/Mathlib/Data/Set/Pointwise/Interval.lean index 88b3b2c811895..f1a9b9d837cec 100644 --- a/Mathlib/Data/Set/Pointwise/Interval.lean +++ b/Mathlib/Data/Set/Pointwise/Interval.lean @@ -304,10 +304,10 @@ theorem preimage_const_sub_Ioo : (fun x => a - x) ⁻¹' Ioo b c = Ioo (a - c) ( -/ --- @[simp] -- Porting note (#10618): simp can prove this modulo `add_comm` +-- simp can prove this modulo `add_comm` theorem image_const_add_Iic : (fun x => a + x) '' Iic b = Iic (a + b) := by simp [add_comm] --- @[simp] -- Porting note (#10618): simp can prove this modulo `add_comm` +-- simp can prove this modulo `add_comm` theorem image_const_add_Iio : (fun x => a + x) '' Iio b = Iio (a + b) := by simp [add_comm] /-! @@ -315,10 +315,8 @@ theorem image_const_add_Iio : (fun x => a + x) '' Iio b = Iio (a + b) := by simp -/ --- @[simp] -- Porting note (#10618): simp can prove this theorem image_add_const_Iic : (fun x => x + a) '' Iic b = Iic (b + a) := by simp --- @[simp] -- Porting note (#10618): simp can prove this theorem image_add_const_Iio : (fun x => x + a) '' Iio b = Iio (b + a) := by simp /-! @@ -459,10 +457,9 @@ theorem preimage_const_sub_uIcc : (fun x => a - x) ⁻¹' [[b, c]] = [[a - b, a simp_rw [← Icc_min_max, preimage_const_sub_Icc] simp only [sub_eq_add_neg, min_add_add_left, max_add_add_left, min_neg_neg, max_neg_neg] --- @[simp] -- Porting note (#10618): simp can prove this module `add_comm` +-- simp can prove this modulo `add_comm` theorem image_const_add_uIcc : (fun x => a + x) '' [[b, c]] = [[a + b, a + c]] := by simp [add_comm] --- @[simp] -- Porting note (#10618): simp can prove this theorem image_add_const_uIcc : (fun x => x + a) '' [[b, c]] = [[b + a, c + a]] := by simp @[simp] diff --git a/Mathlib/Data/Sign.lean b/Mathlib/Data/Sign.lean index a492a43728d98..e5621d92ed798 100644 --- a/Mathlib/Data/Sign.lean +++ b/Mathlib/Data/Sign.lean @@ -375,7 +375,6 @@ section OrderedSemiring variable [OrderedSemiring α] [DecidableRel ((· < ·) : α → α → Prop)] [Nontrivial α] --- @[simp] -- Porting note (#10618): simp can prove this theorem sign_one : sign (1 : α) = 1 := sign_pos zero_lt_one diff --git a/Mathlib/Data/Sym/Basic.lean b/Mathlib/Data/Sym/Basic.lean index 6c89905364aa0..8d71b72273208 100644 --- a/Mathlib/Data/Sym/Basic.lean +++ b/Mathlib/Data/Sym/Basic.lean @@ -176,7 +176,6 @@ theorem mem_coe : a ∈ (s : Multiset α) ↔ a ∈ s := theorem mem_cons_of_mem (h : a ∈ s) : a ∈ b ::ₛ s := Multiset.mem_cons_of_mem h ---@[simp] Porting note (#10618): simp can prove it theorem mem_cons_self (a : α) (s : Sym α n) : a ∈ a ::ₛ s := Multiset.mem_cons_self a s.1 diff --git a/Mathlib/Data/Vector/Basic.lean b/Mathlib/Data/Vector/Basic.lean index 5e7c6ea9c6715..50418d58cb02f 100644 --- a/Mathlib/Data/Vector/Basic.lean +++ b/Mathlib/Data/Vector/Basic.lean @@ -214,7 +214,6 @@ theorem get_zero : ∀ v : Vector α n.succ, get v 0 = head v theorem head_ofFn {n : ℕ} (f : Fin n.succ → α) : head (ofFn f) = f 0 := by rw [← get_zero, get_ofFn] ---@[simp] Porting note (#10618): simp can prove it theorem get_cons_zero (a : α) (v : Vector α n) : get (a ::ᵥ v) 0 = a := by simp [get_zero] /-- Accessing the nth element of a vector made up diff --git a/Mathlib/Data/ZMod/Basic.lean b/Mathlib/Data/ZMod/Basic.lean index b8c5302f15179..cfb7aa79acee7 100644 --- a/Mathlib/Data/ZMod/Basic.lean +++ b/Mathlib/Data/ZMod/Basic.lean @@ -129,7 +129,6 @@ theorem ringChar_zmod_n (n : ℕ) : ringChar (ZMod n) = n := by rw [ringChar.eq_iff] exact ZMod.charP n --- @[simp] -- Porting note (#10618): simp can prove this theorem natCast_self (n : ℕ) : (n : ZMod n) = 0 := CharP.cast_eq_zero (ZMod n) n From dff41254dfecc8d01994c140927452e38ef3bbd9 Mon Sep 17 00:00:00 2001 From: Dagur Asgeirsson Date: Tue, 22 Oct 2024 19:54:05 +0000 Subject: [PATCH 301/505] chore(CategoryTheory/Sites): add tests for sheafification instances (#17815) This PR adds a few tests to make sure that sheafification exists for sheaves of sets and sheaves of modules, under different size restrictions on the defining site. --- .../Sites/ConcreteSheafification.lean | 34 +++++++++++++++++++ 1 file changed, 34 insertions(+) create mode 100644 test/CategoryTheory/Sites/ConcreteSheafification.lean diff --git a/test/CategoryTheory/Sites/ConcreteSheafification.lean b/test/CategoryTheory/Sites/ConcreteSheafification.lean new file mode 100644 index 0000000000000..8f5491183bbc6 --- /dev/null +++ b/test/CategoryTheory/Sites/ConcreteSheafification.lean @@ -0,0 +1,34 @@ +import Mathlib.Algebra.Category.ModuleCat.Colimits +import Mathlib.Algebra.Category.ModuleCat.Limits +import Mathlib.Algebra.Category.ModuleCat.FilteredColimits +import Mathlib.CategoryTheory.Sites.Equivalence + +universe u + +open CategoryTheory + +section Small + +variable {C : Type u} [SmallCategory C] (J : GrothendieckTopology C) + +example : HasSheafify J (Type u) := inferInstance + +example (R : Type u) [Ring R] : HasSheafify J (ModuleCat.{u} R) := inferInstance + +end Small + +section Large + +variable {C : Type (u+1)} [LargeCategory C] (J : GrothendieckTopology C) + +example : HasSheafify J (Type (u+1)) := inferInstance + +example (R : Type (u+1)) [Ring R] : HasSheafify J (ModuleCat.{u+1} R) := inferInstance + +variable [EssentiallySmall.{u} C] + +example : HasSheafify J (Type u) := inferInstance + +example (R : Type u) [Ring R] : HasSheafify J (ModuleCat.{u} R) := inferInstance + +end Large From 42d44d43b9d4ce35d079ecee2ecd173eaf640091 Mon Sep 17 00:00:00 2001 From: Pietro Monticone <38562595+pitmonticone@users.noreply.github.com> Date: Tue, 22 Oct 2024 19:54:06 +0000 Subject: [PATCH 302/505] feat(Data/Finset/Basic): add instances (#17976) Co-authored-by: @Command-Master Upstreamed from the [EquationalTheories](https://github.com/teorth/equational_theories) project. - [x] depends on: #17971 --- Mathlib/Data/Finset/Basic.lean | 18 +++++++++++++++--- 1 file changed, 15 insertions(+), 3 deletions(-) diff --git a/Mathlib/Data/Finset/Basic.lean b/Mathlib/Data/Finset/Basic.lean index 8dad68581aed6..7f0708bdd080c 100644 --- a/Mathlib/Data/Finset/Basic.lean +++ b/Mathlib/Data/Finset/Basic.lean @@ -921,10 +921,9 @@ theorem disjUnion_singleton (s : Finset α) (a : α) (h : Disjoint s {a}) : /-! ### insert -/ - section Insert -variable [DecidableEq α] {s t u v : Finset α} {a b : α} +variable [DecidableEq α] {s t u v : Finset α} {a b : α} {f : α → β} /-- `insert a s` is the set `{a} ∪ s` containing `a` and the elements of `s`. -/ instance : Insert α (Finset α) := @@ -2766,7 +2765,8 @@ end Finset namespace List -variable [DecidableEq α] {l l' : List α} {a : α} +variable [DecidableEq α] {l l' : List α} {a : α} {f : α → β} + {s : Finset α} {t : Set β} {t' : Finset β} /-- `toFinset l` removes duplicates from the list `l` to produce a finset. -/ def toFinset (l : List α) : Finset α := @@ -2807,6 +2807,18 @@ theorem toFinset_surjective : Surjective (toFinset : List α → Finset α) := f let ⟨l, _, hls⟩ := toFinset_surj_on (Set.mem_univ s) ⟨l, hls⟩ +instance [DecidableEq β] : Decidable (Set.SurjOn f s t') := + inferInstanceAs (Decidable (∀ x ∈ t', ∃ y ∈ s, f y = x)) + +instance [DecidableEq β] : Decidable (Set.InjOn f s) := + inferInstanceAs (Decidable (∀ x ∈ s, ∀ y ∈ s, f x = f y → x = y)) + +instance [DecidablePred (· ∈ t)] : Decidable (Set.MapsTo f s t) := + inferInstanceAs (Decidable (∀ x ∈ s, f x ∈ t)) + +instance [DecidableEq β] : Decidable (Set.BijOn f s t') := + inferInstanceAs (Decidable (_ ∧ _ ∧ _)) + theorem toFinset_eq_iff_perm_dedup : l.toFinset = l'.toFinset ↔ l.dedup ~ l'.dedup := by simp [Finset.ext_iff, perm_ext_iff_of_nodup (nodup_dedup _) (nodup_dedup _)] From 054a307e3c9b34d48f912933960c46ae7afdeb47 Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Tue, 22 Oct 2024 20:35:58 +0000 Subject: [PATCH 303/505] chore: fix two module docstrings (#18038) --- Mathlib/Data/DFinsupp/Basic.lean | 2 +- Mathlib/LinearAlgebra/DFinsupp.lean | 9 +++++---- 2 files changed, 6 insertions(+), 5 deletions(-) diff --git a/Mathlib/Data/DFinsupp/Basic.lean b/Mathlib/Data/DFinsupp/Basic.lean index df567e316a5ee..9a705f9e2e263 100644 --- a/Mathlib/Data/DFinsupp/Basic.lean +++ b/Mathlib/Data/DFinsupp/Basic.lean @@ -17,7 +17,7 @@ import Mathlib.Order.ConditionallyCompleteLattice.Basic /-! # Dependent functions with finite support -For a non-dependent version see `data/finsupp.lean`. +For a non-dependent version see `Mathlib.Data.Finsupp.Defs`. ## Notation diff --git a/Mathlib/LinearAlgebra/DFinsupp.lean b/Mathlib/LinearAlgebra/DFinsupp.lean index 52dc2d02af921..5da7fbab05ab4 100644 --- a/Mathlib/LinearAlgebra/DFinsupp.lean +++ b/Mathlib/LinearAlgebra/DFinsupp.lean @@ -11,17 +11,17 @@ import Mathlib.LinearAlgebra.LinearIndependent # Properties of the module `Π₀ i, M i` Given an indexed collection of `R`-modules `M i`, the `R`-module structure on `Π₀ i, M i` -is defined in `Data.DFinsupp`. +is defined in `Mathlib.Data.DFinsupp.Basic`. In this file we define `LinearMap` versions of various maps: * `DFinsupp.lsingle a : M →ₗ[R] Π₀ i, M i`: `DFinsupp.single a` as a linear map; -* `DFinsupp.lmk s : (Π i : (↑s : Set ι), M i) →ₗ[R] Π₀ i, M i`: `DFinsupp.single a` as a linear map; +* `DFinsupp.lmk s : (Π i : (↑s : Set ι), M i) →ₗ[R] Π₀ i, M i`: `DFinsupp.mk` as a linear map; * `DFinsupp.lapply i : (Π₀ i, M i) →ₗ[R] M`: the map `fun f ↦ f i` as a linear map; -* `DFinsupp.lsum`: `DFinsupp.sum` or `DFinsupp.liftAddHom` as a `LinearMap`; +* `DFinsupp.lsum`: `DFinsupp.sum` or `DFinsupp.liftAddHom` as a `LinearMap`. ## Implementation notes @@ -62,7 +62,8 @@ theorem lhom_ext ⦃φ ψ : (Π₀ i, M i) →ₗ[R] N⦄ (h : ∀ i x, φ (sing /-- Two `R`-linear maps from `Π₀ i, M i` which agree on each `single i x` agree everywhere. See note [partially-applied ext lemmas]. -After apply this lemma, if `M = R` then it suffices to verify `φ (single a 1) = ψ (single a 1)`. -/ +After applying this lemma, if `M = R` then it suffices to verify +`φ (single a 1) = ψ (single a 1)`. -/ @[ext 1100] theorem lhom_ext' ⦃φ ψ : (Π₀ i, M i) →ₗ[R] N⦄ (h : ∀ i, φ.comp (lsingle i) = ψ.comp (lsingle i)) : φ = ψ := From d50b9137dab94370c198303a58b87846530b7535 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 22 Oct 2024 21:02:22 +0000 Subject: [PATCH 304/505] chore(Data): use newly introduced finset notation (#18059) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit For `s : Finset α`, replace `s.card` with `#s`, `s.filter p` with `{x ∈ s | p x}`, `∑ x ∈ s.filter p, f x` with `∑ x ∈ s with p x, f x`. Rewrap lines around and open the `Finset` locale where necessary. --- Mathlib/Data/Complex/Exponential.lean | 18 +++-- Mathlib/Data/DFinsupp/Basic.lean | 4 +- Mathlib/Data/DFinsupp/Interval.lean | 19 +++--- Mathlib/Data/Finset/Finsupp.lean | 5 +- Mathlib/Data/Finset/NAry.lean | 18 ++--- Mathlib/Data/Finset/Slice.lean | 16 ++--- Mathlib/Data/Finset/Sups.lean | 28 +++----- Mathlib/Data/Finsupp/Basic.lean | 2 +- Mathlib/Data/Finsupp/Defs.lean | 14 ++-- Mathlib/Data/Finsupp/Indicator.lean | 2 +- Mathlib/Data/Finsupp/Interval.lean | 14 ++-- Mathlib/Data/Fintype/BigOperators.lean | 26 ++++--- Mathlib/Data/Fintype/Card.lean | 83 +++++++++++------------ Mathlib/Data/Fintype/Fin.lean | 11 ++- Mathlib/Data/Fintype/Perm.lean | 2 +- Mathlib/Data/Fintype/Pi.lean | 8 +-- Mathlib/Data/Fintype/Powerset.lean | 10 +-- Mathlib/Data/Fintype/Sort.lean | 6 +- Mathlib/Data/Fintype/Sum.lean | 4 +- Mathlib/Data/Int/CardIntervalMod.lean | 52 +++++++------- Mathlib/Data/Int/Interval.lean | 18 ++--- Mathlib/Data/Multiset/Interval.lean | 8 +-- Mathlib/Data/Nat/Choose/Sum.lean | 8 +-- Mathlib/Data/Nat/Count.lean | 10 +-- Mathlib/Data/Nat/Factorization/Basic.lean | 15 ++-- Mathlib/Data/Nat/Factorization/Defs.lean | 2 +- Mathlib/Data/Nat/Multiplicity.lean | 24 +++---- Mathlib/Data/Nat/Nth.lean | 66 +++++++++--------- Mathlib/Data/Nat/Squarefree.lean | 8 +-- Mathlib/Data/Nat/Totient.lean | 33 +++++---- Mathlib/Data/PNat/Interval.lean | 10 +-- Mathlib/Data/Pi/Interval.lean | 24 +++---- Mathlib/Data/Sigma/Interval.lean | 8 +-- Mathlib/Data/Sign.lean | 2 +- Mathlib/Data/Sym/Card.lean | 14 ++-- Mathlib/Data/Sym/Sym2.lean | 5 +- 36 files changed, 283 insertions(+), 314 deletions(-) diff --git a/Mathlib/Data/Complex/Exponential.lean b/Mathlib/Data/Complex/Exponential.lean index e3fd8a0bac2f2..e5502836c9e2e 100644 --- a/Mathlib/Data/Complex/Exponential.lean +++ b/Mathlib/Data/Complex/Exponential.lean @@ -1061,10 +1061,9 @@ end Real namespace Complex theorem sum_div_factorial_le {α : Type*} [LinearOrderedField α] (n j : ℕ) (hn : 0 < n) : - (∑ m ∈ filter (fun k => n ≤ k) (range j), - (1 / m.factorial : α)) ≤ n.succ / (n.factorial * n) := + (∑ m ∈ range j with n ≤ m, (1 / m.factorial : α)) ≤ n.succ / (n.factorial * n) := calc - (∑ m ∈ filter (fun k => n ≤ k) (range j), (1 / m.factorial : α)) = + (∑ m ∈ range j with n ≤ m, (1 / m.factorial : α)) = ∑ m ∈ range (j - n), (1 / ((m + n).factorial : α)) := by refine sum_nbij' (· - n) (· + n) ?_ ?_ ?_ ?_ ?_ <;> simp (config := { contextual := true }) [lt_tsub_iff_right, tsub_add_cancel_of_le] @@ -1098,20 +1097,19 @@ theorem exp_bound {x : ℂ} (hx : abs x ≤ 1) {n : ℕ} (hn : 0 < n) : abs x ^ n * ((n.succ : ℝ) * (n.factorial * n : ℝ)⁻¹) rw [sum_range_sub_sum_range hj] calc - abs (∑ m ∈ (range j).filter fun k => n ≤ k, (x ^ m / m.factorial : ℂ)) = - abs (∑ m ∈ (range j).filter fun k => n ≤ k, - (x ^ n * (x ^ (m - n) / m.factorial) : ℂ)) := by + abs (∑ m ∈ range j with n ≤ m, (x ^ m / m.factorial : ℂ)) + = abs (∑ m ∈ range j with n ≤ m, (x ^ n * (x ^ (m - n) / m.factorial) : ℂ)) := by refine congr_arg abs (sum_congr rfl fun m hm => ?_) rw [mem_filter, mem_range] at hm rw [← mul_div_assoc, ← pow_add, add_tsub_cancel_of_le hm.2] - _ ≤ ∑ m ∈ filter (fun k => n ≤ k) (range j), abs (x ^ n * (x ^ (m - n) / m.factorial)) := - (IsAbsoluteValue.abv_sum Complex.abs _ _) - _ ≤ ∑ m ∈ filter (fun k => n ≤ k) (range j), abs x ^ n * (1 / m.factorial) := by + _ ≤ ∑ m ∈ range j with n ≤ m, abs (x ^ n * (x ^ (m - n) / m.factorial)) := + IsAbsoluteValue.abv_sum Complex.abs .. + _ ≤ ∑ m ∈ range j with n ≤ m, abs x ^ n * (1 / m.factorial) := by simp_rw [map_mul, map_pow, map_div₀, abs_natCast] gcongr rw [abv_pow abs] exact pow_le_one₀ (abs.nonneg _) hx - _ = abs x ^ n * ∑ m ∈ (range j).filter fun k => n ≤ k, (1 / m.factorial : ℝ) := by + _ = abs x ^ n * ∑ m ∈ range j with n ≤ m, (1 / m.factorial : ℝ) := by simp [abs_mul, abv_pow abs, abs_div, ← mul_sum] _ ≤ abs x ^ n * (n.succ * (n.factorial * n : ℝ)⁻¹) := by gcongr diff --git a/Mathlib/Data/DFinsupp/Basic.lean b/Mathlib/Data/DFinsupp/Basic.lean index 9a705f9e2e263..c08ce66284b1d 100644 --- a/Mathlib/Data/DFinsupp/Basic.lean +++ b/Mathlib/Data/DFinsupp/Basic.lean @@ -1118,7 +1118,7 @@ theorem filter_def (f : Π₀ i, β i) : f.filter p = mk (f.support.filter p) fu ext i; by_cases h1 : p i <;> by_cases h2 : f i ≠ 0 <;> simp at h2 <;> simp [h1, h2] @[simp] -theorem support_filter (f : Π₀ i, β i) : (f.filter p).support = f.support.filter p := by +theorem support_filter (f : Π₀ i, β i) : (f.filter p).support = {x ∈ f.support | p x} := by ext i; by_cases h : p i <;> simp [h] theorem subtypeDomain_def (f : Π₀ i, β i) : @@ -1715,7 +1715,7 @@ theorem sumAddHom_comp_single [∀ i, AddZeroClass (β i)] [AddCommMonoid γ] (f theorem sumAddHom_apply [∀ i, AddZeroClass (β i)] [∀ (i) (x : β i), Decidable (x ≠ 0)] [AddCommMonoid γ] (φ : ∀ i, β i →+ γ) (f : Π₀ i, β i) : sumAddHom φ f = f.sum fun x => φ x := by rcases f with ⟨f, s, hf⟩ - change (∑ i ∈ _, _) = ∑ i ∈ Finset.filter _ _, _ + change (∑ i ∈ _, _) = ∑ i ∈ _ with _, _ rw [Finset.sum_filter, Finset.sum_congr rfl] intro i _ dsimp only [coe_mk', Subtype.coe_mk] at * diff --git a/Mathlib/Data/DFinsupp/Interval.lean b/Mathlib/Data/DFinsupp/Interval.lean index 73f9fad8750ec..8cb45f7f31389 100644 --- a/Mathlib/Data/DFinsupp/Interval.lean +++ b/Mathlib/Data/DFinsupp/Interval.lean @@ -35,8 +35,7 @@ def dfinsupp (s : Finset ι) (t : ∀ i, Finset (α i)) : Finset (Π₀ i, α i) convert congr_fun h ⟨i, hi⟩⟩ @[simp] -theorem card_dfinsupp (s : Finset ι) (t : ∀ i, Finset (α i)) : - (s.dfinsupp t).card = ∏ i ∈ s, (t i).card := +theorem card_dfinsupp (s : Finset ι) (t : ∀ i, Finset (α i)) : #(s.dfinsupp t) = ∏ i ∈ s, #(t i) := (card_map _).trans <| card_pi _ _ variable [∀ i, DecidableEq (α i)] @@ -134,7 +133,7 @@ theorem mem_pi {f : Π₀ i, Finset (α i)} {g : Π₀ i, α i} : g ∈ f.pi ↔ mem_dfinsupp_iff_of_support_subset <| Subset.refl _ @[simp] -theorem card_pi (f : Π₀ i, Finset (α i)) : f.pi.card = f.prod fun i => (f i).card := by +theorem card_pi (f : Π₀ i, Finset (α i)) : #f.pi = f.prod fun i ↦ #(f i) := by rw [pi, card_dfinsupp] exact Finset.prod_congr rfl fun i _ => by simp only [Pi.natCast_apply, Nat.cast_id] @@ -157,16 +156,16 @@ variable (f g : Π₀ i, α i) theorem Icc_eq : Icc f g = (f.support ∪ g.support).dfinsupp (f.rangeIcc g) := rfl -theorem card_Icc : (Icc f g).card = ∏ i ∈ f.support ∪ g.support, (Icc (f i) (g i)).card := +lemma card_Icc : #(Icc f g) = ∏ i ∈ f.support ∪ g.support, #(Icc (f i) (g i)) := card_dfinsupp _ _ -theorem card_Ico : (Ico f g).card = (∏ i ∈ f.support ∪ g.support, (Icc (f i) (g i)).card) - 1 := by +lemma card_Ico : #(Ico f g) = (∏ i ∈ f.support ∪ g.support, #(Icc (f i) (g i))) - 1 := by rw [card_Ico_eq_card_Icc_sub_one, card_Icc] -theorem card_Ioc : (Ioc f g).card = (∏ i ∈ f.support ∪ g.support, (Icc (f i) (g i)).card) - 1 := by +lemma card_Ioc : #(Ioc f g) = (∏ i ∈ f.support ∪ g.support, #(Icc (f i) (g i))) - 1 := by rw [card_Ioc_eq_card_Icc_sub_one, card_Icc] -theorem card_Ioo : (Ioo f g).card = (∏ i ∈ f.support ∪ g.support, (Icc (f i) (g i)).card) - 2 := by +lemma card_Ioo : #(Ioo f g) = (∏ i ∈ f.support ∪ g.support, #(Icc (f i) (g i))) - 2 := by rw [card_Ioo_eq_card_Icc_sub_two, card_Icc] end PartialOrder @@ -175,7 +174,7 @@ section Lattice variable [DecidableEq ι] [∀ i, DecidableEq (α i)] [∀ i, Lattice (α i)] [∀ i, Zero (α i)] [∀ i, LocallyFiniteOrder (α i)] (f g : Π₀ i, α i) -theorem card_uIcc : (uIcc f g).card = ∏ i ∈ f.support ∪ g.support, (uIcc (f i) (g i)).card := by +lemma card_uIcc : #(uIcc f g) = ∏ i ∈ f.support ∪ g.support, #(uIcc (f i) (g i)) := by rw [← support_inf_union_support_sup]; exact card_Icc _ _ end Lattice @@ -186,11 +185,11 @@ variable [DecidableEq ι] [∀ i, DecidableEq (α i)] variable [∀ i, CanonicallyOrderedAddCommMonoid (α i)] [∀ i, LocallyFiniteOrder (α i)] variable (f : Π₀ i, α i) -theorem card_Iic : (Iic f).card = ∏ i ∈ f.support, (Iic (f i)).card := by +lemma card_Iic : #(Iic f) = ∏ i ∈ f.support, #(Iic (f i)) := by simp_rw [Iic_eq_Icc, card_Icc, DFinsupp.bot_eq_zero, support_zero, empty_union, zero_apply, bot_eq_zero] -theorem card_Iio : (Iio f).card = (∏ i ∈ f.support, (Iic (f i)).card) - 1 := by +lemma card_Iio : #(Iio f) = (∏ i ∈ f.support, #(Iic (f i))) - 1 := by rw [card_Iio_eq_card_Iic_sub_one, card_Iic] end CanonicallyOrdered diff --git a/Mathlib/Data/Finset/Finsupp.lean b/Mathlib/Data/Finset/Finsupp.lean index 3d7111a8e4605..43e04648e3ccf 100644 --- a/Mathlib/Data/Finset/Finsupp.lean +++ b/Mathlib/Data/Finset/Finsupp.lean @@ -70,8 +70,7 @@ theorem mem_finsupp_iff_of_support_subset {t : ι →₀ Finset α} (ht : t.supp · rwa [H, mem_zero] at h @[simp] -theorem card_finsupp (s : Finset ι) (t : ι → Finset α) : - (s.finsupp t).card = ∏ i ∈ s, (t i).card := +theorem card_finsupp (s : Finset ι) (t : ι → Finset α) : #(s.finsupp t) = ∏ i ∈ s, #(t i) := (card_map _).trans <| card_pi _ _ end Finset @@ -90,7 +89,7 @@ theorem mem_pi {f : ι →₀ Finset α} {g : ι →₀ α} : g ∈ f.pi ↔ ∀ mem_finsupp_iff_of_support_subset <| Subset.refl _ @[simp] -theorem card_pi (f : ι →₀ Finset α) : f.pi.card = f.prod fun i => (f i).card := by +theorem card_pi (f : ι →₀ Finset α) : #f.pi = f.prod fun i ↦ #(f i) := by rw [pi, card_finsupp] exact Finset.prod_congr rfl fun i _ => by simp only [Pi.natCast_apply, Nat.cast_id] diff --git a/Mathlib/Data/Finset/NAry.lean b/Mathlib/Data/Finset/NAry.lean index db73633e28a15..a7d6db38c1b90 100644 --- a/Mathlib/Data/Finset/NAry.lean +++ b/Mathlib/Data/Finset/NAry.lean @@ -46,16 +46,16 @@ theorem coe_image₂ (f : α → β → γ) (s : Finset α) (t : Finset β) : Set.ext fun _ => mem_image₂ theorem card_image₂_le (f : α → β → γ) (s : Finset α) (t : Finset β) : - (image₂ f s t).card ≤ s.card * t.card := + #(image₂ f s t) ≤ #s * #t := card_image_le.trans_eq <| card_product _ _ theorem card_image₂_iff : - (image₂ f s t).card = s.card * t.card ↔ (s ×ˢ t : Set (α × β)).InjOn fun x => f x.1 x.2 := by + #(image₂ f s t) = #s * #t ↔ (s ×ˢ t : Set (α × β)).InjOn fun x => f x.1 x.2 := by rw [← card_product, ← coe_product] exact card_image_iff theorem card_image₂ (hf : Injective2 f) (s : Finset α) (t : Finset β) : - (image₂ f s t).card = s.card * t.card := + #(image₂ f s t) = #s * #t := (card_image_of_injective _ hf.uncurry).trans <| card_product _ _ theorem mem_image₂_of_mem (ha : a ∈ s) (hb : b ∈ t) : f a b ∈ image₂ f s t := @@ -196,11 +196,11 @@ theorem image₂_congr' (h : ∀ a b, f a b = f' a b) : image₂ f s t = image variable (s t) -theorem card_image₂_singleton_left (hf : Injective (f a)) : (image₂ f {a} t).card = t.card := by +theorem card_image₂_singleton_left (hf : Injective (f a)) : #(image₂ f {a} t) = #t := by rw [image₂_singleton_left, card_image_of_injective _ hf] theorem card_image₂_singleton_right (hf : Injective fun a => f a b) : - (image₂ f s {b}).card = s.card := by rw [image₂_singleton_right, card_image_of_injective _ hf] + #(image₂ f s {b}) = #s := by rw [image₂_singleton_right, card_image_of_injective _ hf] theorem image₂_singleton_inter [DecidableEq β] (t₁ t₂ : Finset β) (hf : Injective (f a)) : image₂ f {a} (t₁ ∩ t₂) = image₂ f {a} t₁ ∩ image₂ f {a} t₂ := by @@ -211,13 +211,13 @@ theorem image₂_inter_singleton [DecidableEq α] (s₁ s₂ : Finset α) (hf : simp_rw [image₂_singleton_right, image_inter _ _ hf] theorem card_le_card_image₂_left {s : Finset α} (hs : s.Nonempty) (hf : ∀ a, Injective (f a)) : - t.card ≤ (image₂ f s t).card := by + #t ≤ #(image₂ f s t) := by obtain ⟨a, ha⟩ := hs rw [← card_image₂_singleton_left _ (hf a)] exact card_le_card (image₂_subset_right <| singleton_subset_iff.2 ha) theorem card_le_card_image₂_right {t : Finset β} (ht : t.Nonempty) - (hf : ∀ b, Injective fun a => f a b) : s.card ≤ (image₂ f s t).card := by + (hf : ∀ b, Injective fun a => f a b) : #s ≤ #(image₂ f s t) := by obtain ⟨b, hb⟩ := ht rw [← card_image₂_singleton_right _ (hf b)] exact card_le_card (image₂_subset_left <| singleton_subset_iff.2 hb) @@ -431,7 +431,7 @@ theorem image₂_right_identity {f : γ → β → γ} {b : β} (h : ∀ a, f a applications are disjoint (but not necessarily distinct!), then the size of `t` divides the size of `Finset.image₂ f s t`. -/ theorem card_dvd_card_image₂_right (hf : ∀ a ∈ s, Injective (f a)) - (hs : ((fun a => t.image <| f a) '' s).PairwiseDisjoint id) : t.card ∣ (image₂ f s t).card := by + (hs : ((fun a => t.image <| f a) '' s).PairwiseDisjoint id) : #t ∣ #(image₂ f s t) := by classical induction' s using Finset.induction with a s _ ih · simp @@ -453,7 +453,7 @@ applications are disjoint (but not necessarily distinct!), then the size of `s` `Finset.image₂ f s t`. -/ theorem card_dvd_card_image₂_left (hf : ∀ b ∈ t, Injective fun a => f a b) (ht : ((fun b => s.image fun a => f a b) '' t).PairwiseDisjoint id) : - s.card ∣ (image₂ f s t).card := by rw [← image₂_swap]; exact card_dvd_card_image₂_right hf ht + #s ∣ #(image₂ f s t) := by rw [← image₂_swap]; exact card_dvd_card_image₂_right hf ht /-- 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. -/ diff --git a/Mathlib/Data/Finset/Slice.lean b/Mathlib/Data/Finset/Slice.lean index 9ad5073eabffc..577c05063d19a 100644 --- a/Mathlib/Data/Finset/Slice.lean +++ b/Mathlib/Data/Finset/Slice.lean @@ -39,13 +39,12 @@ variable {A B : Set (Finset α)} {s : Finset α} {r : ℕ} /-- `Sized r A` means that every Finset in `A` has size `r`. -/ -def Sized (r : ℕ) (A : Set (Finset α)) : Prop := - ∀ ⦃x⦄, x ∈ A → card x = r +def Sized (r : ℕ) (A : Set (Finset α)) : Prop := ∀ ⦃x⦄, x ∈ A → #x = r theorem Sized.mono (h : A ⊆ B) (hB : B.Sized r) : A.Sized r := fun _x hx => hB <| h hx @[simp] lemma sized_empty : (∅ : Set (Finset α)).Sized r := by simp [Sized] -@[simp] lemma sized_singleton : ({s} : Set (Finset α)).Sized r ↔ s.card = r := by simp [Sized] +@[simp] lemma sized_singleton : ({s} : Set (Finset α)).Sized r ↔ #s = r := by simp [Sized] theorem sized_union : (A ∪ B).Sized r ↔ A.Sized r ∧ B.Sized r := ⟨fun hA => ⟨hA.mono subset_union_left, hA.mono subset_union_right⟩, fun hA _x hx => @@ -96,7 +95,7 @@ theorem subset_powersetCard_univ_iff : 𝒜 ⊆ powersetCard r univ ↔ (𝒜 : alias ⟨_, _root_.Set.Sized.subset_powersetCard_univ⟩ := subset_powersetCard_univ_iff theorem _root_.Set.Sized.card_le (h𝒜 : (𝒜 : Set (Finset α)).Sized r) : - card 𝒜 ≤ (Fintype.card α).choose r := by + #𝒜 ≤ (Fintype.card α).choose r := by rw [Fintype.card, ← card_powersetCard] exact card_le_card (subset_powersetCard_univ_iff.mpr h𝒜) @@ -110,15 +109,14 @@ section Slice variable {𝒜 : Finset (Finset α)} {A A₁ A₂ : Finset α} {r r₁ r₂ : ℕ} /-- The `r`-th slice of a set family is the subset of its elements which have cardinality `r`. -/ -def slice (𝒜 : Finset (Finset α)) (r : ℕ) : Finset (Finset α) := - 𝒜.filter fun i => i.card = r +def slice (𝒜 : Finset (Finset α)) (r : ℕ) : Finset (Finset α) := {A ∈ 𝒜 | #A = r} -- Porting note: old code: scoped[FinsetFamily] @[inherit_doc] scoped[Finset] infixl:90 " # " => Finset.slice /-- `A` is in the `r`-th slice of `𝒜` iff it's in `𝒜` and has cardinality `r`. -/ -theorem mem_slice : A ∈ 𝒜 # r ↔ A ∈ 𝒜 ∧ A.card = r := +theorem mem_slice : A ∈ 𝒜 # r ↔ A ∈ 𝒜 ∧ #A = r := mem_filter /-- The `r`-th slice of `𝒜` is a subset of `𝒜`. -/ @@ -143,10 +141,10 @@ variable [Fintype α] (𝒜) @[simp] theorem biUnion_slice [DecidableEq α] : (Iic <| Fintype.card α).biUnion 𝒜.slice = 𝒜 := Subset.antisymm (biUnion_subset.2 fun _r _ => slice_subset) fun s hs => - mem_biUnion.2 ⟨s.card, mem_Iic.2 <| s.card_le_univ, mem_slice.2 <| ⟨hs, rfl⟩⟩ + mem_biUnion.2 ⟨#s, mem_Iic.2 <| s.card_le_univ, mem_slice.2 <| ⟨hs, rfl⟩⟩ @[simp] -theorem sum_card_slice : (∑ r ∈ Iic (Fintype.card α), (𝒜 # r).card) = 𝒜.card := by +theorem sum_card_slice : ∑ r ∈ Iic (Fintype.card α), #(𝒜 # r) = #𝒜 := by letI := Classical.decEq α rw [← card_biUnion, biUnion_slice] exact Finset.pairwiseDisjoint_slice.subset (Set.subset_univ _) diff --git a/Mathlib/Data/Finset/Sups.lean b/Mathlib/Data/Finset/Sups.lean index 5188385616bb7..15185f0898be6 100644 --- a/Mathlib/Data/Finset/Sups.lean +++ b/Mathlib/Data/Finset/Sups.lean @@ -67,11 +67,9 @@ variable (s t) theorem coe_sups : (↑(s ⊻ t) : Set α) = ↑s ⊻ ↑t := coe_image₂ _ _ _ -theorem card_sups_le : (s ⊻ t).card ≤ s.card * t.card := - card_image₂_le _ _ _ +theorem card_sups_le : #(s ⊻ t) ≤ #s * #t := card_image₂_le _ _ _ -theorem card_sups_iff : - (s ⊻ t).card = s.card * t.card ↔ (s ×ˢ t : Set (α × α)).InjOn fun x => x.1 ⊔ x.2 := +theorem card_sups_iff : #(s ⊻ t) = #s * #t ↔ (s ×ˢ t : Set (α × α)).InjOn fun x => x.1 ⊔ x.2 := card_image₂_iff variable {s s₁ s₂ t t₁ t₂ u} @@ -162,7 +160,7 @@ lemma sups_subset_self : s ⊻ s ⊆ s ↔ SupClosed (s : Set α) := sups_subset @[simp] lemma univ_sups_univ [Fintype α] : (univ : Finset α) ⊻ univ = univ := by simp lemma filter_sups_le [@DecidableRel α (· ≤ ·)] (s t : Finset α) (a : α) : - (s ⊻ t).filter (· ≤ a) = s.filter (· ≤ a) ⊻ t.filter (· ≤ a) := by + {b ∈ s ⊻ t | b ≤ a} = {b ∈ s | b ≤ a} ⊻ {b ∈ t | b ≤ a} := by simp only [← coe_inj, coe_filter, coe_sups, ← mem_coe, Set.sep_sups_le] variable (s t u) @@ -214,11 +212,9 @@ variable (s t) theorem coe_infs : (↑(s ⊼ t) : Set α) = ↑s ⊼ ↑t := coe_image₂ _ _ _ -theorem card_infs_le : (s ⊼ t).card ≤ s.card * t.card := - card_image₂_le _ _ _ +theorem card_infs_le : #(s ⊼ t) ≤ #s * #t := card_image₂_le _ _ _ -theorem card_infs_iff : - (s ⊼ t).card = s.card * t.card ↔ (s ×ˢ t : Set (α × α)).InjOn fun x => x.1 ⊓ x.2 := +theorem card_infs_iff : #(s ⊼ t) = #s * #t ↔ (s ×ˢ t : Set (α × α)).InjOn fun x => x.1 ⊓ x.2 := card_image₂_iff variable {s s₁ s₂ t t₁ t₂ u} @@ -309,7 +305,7 @@ lemma infs_self_subset : s ⊼ s ⊆ s ↔ InfClosed (s : Set α) := infs_subset @[simp] lemma univ_infs_univ [Fintype α] : (univ : Finset α) ⊼ univ = univ := by simp lemma filter_infs_le [@DecidableRel α (· ≤ ·)] (s t : Finset α) (a : α) : - (s ⊼ t).filter (a ≤ ·) = s.filter (a ≤ ·) ⊼ t.filter (a ≤ ·) := by + {b ∈ s ⊼ t | a ≤ b} = {b ∈ s | a ≤ b} ⊼ {b ∈ t | a ≤ b} := by simp only [← coe_inj, coe_filter, coe_infs, ← mem_coe, Set.sep_infs_le] variable (s t u) @@ -396,8 +392,7 @@ variable [SemilatticeSup α] [OrderBot α] [@DecidableRel α Disjoint] (s s₁ s /-- The finset of elements of the form `a ⊔ b` where `a ∈ s`, `b ∈ t` and `a` and `b` are disjoint. -/ -def disjSups : Finset α := - ((s ×ˢ t).filter fun ab : α × α => Disjoint ab.1 ab.2).image fun ab => ab.1 ⊔ ab.2 +def disjSups : Finset α := {ab ∈ s ×ˢ t | Disjoint ab.1 ab.2}.image fun ab => ab.1 ⊔ ab.2 @[inherit_doc] scoped[FinsetFamily] infixl:74 " ○ " => Finset.disjSups @@ -416,7 +411,7 @@ theorem disjSups_subset_sups : s ○ t ⊆ s ⊻ t := by variable (s t) -theorem card_disjSups_le : (s ○ t).card ≤ s.card * t.card := +theorem card_disjSups_le : #(s ○ t) ≤ #s * #t := (card_le_card disjSups_subset_sups).trans <| card_sups_le _ _ variable {s s₁ s₂ t t₁ t₂} @@ -534,10 +529,9 @@ variable (s t) @[simp, norm_cast] lemma coe_diffs : (↑(s \\ t) : Set α) = Set.image2 (· \ ·) s t := coe_image₂ _ _ _ -lemma card_diffs_le : (s \\ t).card ≤ s.card * t.card := card_image₂_le _ _ _ +lemma card_diffs_le : #(s \\ t) ≤ #s * #t := card_image₂_le _ _ _ -lemma card_diffs_iff : - (s \\ t).card = s.card * t.card ↔ (s ×ˢ t : Set (α × α)).InjOn fun x ↦ x.1 \ x.2 := +lemma card_diffs_iff : #(s \\ t) = #s * #t ↔ (s ×ˢ t : Set (α × α)).InjOn fun x ↦ x.1 \ x.2 := card_image₂_iff variable {s s₁ s₂ t t₁ t₂ u} @@ -619,7 +613,7 @@ variable (s t) @[simp, norm_cast] lemma coe_compls : (↑sᶜˢ : Set α) = compl '' ↑s := coe_map _ _ -@[simp] lemma card_compls : sᶜˢ.card = s.card := card_map _ +@[simp] lemma card_compls : #sᶜˢ = #s := card_map _ variable {s s₁ s₂ t t₁ t₂ u} diff --git a/Mathlib/Data/Finsupp/Basic.lean b/Mathlib/Data/Finsupp/Basic.lean index 9d44b568a65db..e2a9c7d382a55 100644 --- a/Mathlib/Data/Finsupp/Basic.lean +++ b/Mathlib/Data/Finsupp/Basic.lean @@ -799,7 +799,7 @@ theorem filter_apply_pos {a : α} (h : p a) : f.filter p a = f a := if_pos h theorem filter_apply_neg {a : α} (h : ¬p a) : f.filter p a = 0 := if_neg h @[simp] -theorem support_filter : (f.filter p).support = f.support.filter p := rfl +theorem support_filter : (f.filter p).support = {x ∈ f.support | p x} := rfl theorem filter_zero : (0 : α →₀ M).filter p = 0 := by classical rw [← support_eq_empty, support_filter, support_zero, Finset.filter_empty] diff --git a/Mathlib/Data/Finsupp/Defs.lean b/Mathlib/Data/Finsupp/Defs.lean index b44c8ceb8c1bb..31b93c4c9ac8f 100644 --- a/Mathlib/Data/Finsupp/Defs.lean +++ b/Mathlib/Data/Finsupp/Defs.lean @@ -171,7 +171,7 @@ theorem support_eq_empty {f : α →₀ M} : f.support = ∅ ↔ f = 0 := theorem support_nonempty_iff {f : α →₀ M} : f.support.Nonempty ↔ f ≠ 0 := by simp only [Finsupp.support_eq_empty, Finset.nonempty_iff_ne_empty, Ne] -theorem card_support_eq_zero {f : α →₀ M} : card f.support = 0 ↔ f = 0 := by simp +theorem card_support_eq_zero {f : α →₀ M} : #f.support = 0 ↔ f = 0 := by simp instance instDecidableEq [DecidableEq α] [DecidableEq M] : DecidableEq (α →₀ M) := fun f g => decidable_of_iff (f.support = g.support ∧ ∀ a ∈ f.support, f a = g a) ext_iff'.symm @@ -388,11 +388,11 @@ theorem support_eq_singleton' {f : α →₀ M} {a : α} : fun ⟨_b, hb, hf⟩ => hf.symm ▸ support_single_ne_zero _ hb⟩ theorem card_support_eq_one {f : α →₀ M} : - card f.support = 1 ↔ ∃ a, f a ≠ 0 ∧ f = single a (f a) := by + #f.support = 1 ↔ ∃ a, f a ≠ 0 ∧ f = single a (f a) := by simp only [card_eq_one, support_eq_singleton] theorem card_support_eq_one' {f : α →₀ M} : - card f.support = 1 ↔ ∃ a, ∃ b ≠ 0, f = single a b := by + #f.support = 1 ↔ ∃ a, ∃ b ≠ 0, f = single a b := by simp only [card_eq_one, support_eq_singleton'] theorem support_subset_singleton {f : α →₀ M} {a : α} : f.support ⊆ {a} ↔ f = single a (f a) := @@ -403,11 +403,11 @@ theorem support_subset_singleton' {f : α →₀ M} {a : α} : f.support ⊆ {a} rw [hb, support_subset_singleton, single_eq_same]⟩ theorem card_support_le_one [Nonempty α] {f : α →₀ M} : - card f.support ≤ 1 ↔ ∃ a, f = single a (f a) := by + #f.support ≤ 1 ↔ ∃ a, f = single a (f a) := by simp only [card_le_one_iff_subset_singleton, support_subset_singleton] theorem card_support_le_one' [Nonempty α] {f : α →₀ M} : - card f.support ≤ 1 ↔ ∃ a b, f = single a b := by + #f.support ≤ 1 ↔ ∃ a b, f = single a b := by simp only [card_le_one_iff_subset_singleton, support_subset_singleton'] @[simp] @@ -622,7 +622,7 @@ otherwise a better set representation is often available. -/ def onFinset (s : Finset α) (f : α → M) (hf : ∀ a, f a ≠ 0 → a ∈ s) : α →₀ M where support := haveI := Classical.decEq M - s.filter (f · ≠ 0) + {a ∈ s | f a ≠ 0} toFun := f mem_support_toFun := by classical simpa @@ -643,7 +643,7 @@ theorem mem_support_onFinset {s : Finset α} {f : α → M} (hf : ∀ a : α, f theorem support_onFinset [DecidableEq M] {s : Finset α} {f : α → M} (hf : ∀ a : α, f a ≠ 0 → a ∈ s) : - (Finsupp.onFinset s f hf).support = s.filter fun a => f a ≠ 0 := by + (Finsupp.onFinset s f hf).support = {a ∈ s | f a ≠ 0} := by dsimp [onFinset]; congr end OnFinset diff --git a/Mathlib/Data/Finsupp/Indicator.lean b/Mathlib/Data/Finsupp/Indicator.lean index 676bbe7d09324..5c4315a90b1c6 100644 --- a/Mathlib/Data/Finsupp/Indicator.lean +++ b/Mathlib/Data/Finsupp/Indicator.lean @@ -33,7 +33,7 @@ def indicator (s : Finset ι) (f : ∀ i ∈ s, α) : ι →₀ α where if H : i ∈ s then f i H else 0 support := haveI := Classical.decEq α - (s.attach.filter fun i : s => f i.1 i.2 ≠ 0).map (Embedding.subtype _) + ({i | f i.1 i.2 ≠ 0} : Finset s).map (Embedding.subtype _) mem_support_toFun i := by classical simp diff --git a/Mathlib/Data/Finsupp/Interval.lean b/Mathlib/Data/Finsupp/Interval.lean index c0104c16bd32a..9fed894fefb53 100644 --- a/Mathlib/Data/Finsupp/Interval.lean +++ b/Mathlib/Data/Finsupp/Interval.lean @@ -97,19 +97,19 @@ instance instLocallyFiniteOrder : LocallyFiniteOrder (ι →₀ α) := theorem Icc_eq : Icc f g = (f.support ∪ g.support).finsupp (f.rangeIcc g) := rfl -- Porting note: removed [DecidableEq ι] -theorem card_Icc : (Icc f g).card = ∏ i ∈ f.support ∪ g.support, (Icc (f i) (g i)).card := by +theorem card_Icc : #(Icc f g) = ∏ i ∈ f.support ∪ g.support, #(Icc (f i) (g i)):= by simp_rw [Icc_eq, card_finsupp, coe_rangeIcc] -- Porting note: removed [DecidableEq ι] -theorem card_Ico : (Ico f g).card = (∏ i ∈ f.support ∪ g.support, (Icc (f i) (g i)).card) - 1 := by +theorem card_Ico : #(Ico f g) = ∏ i ∈ f.support ∪ g.support, #(Icc (f i) (g i)) - 1 := by rw [card_Ico_eq_card_Icc_sub_one, card_Icc] -- Porting note: removed [DecidableEq ι] -theorem card_Ioc : (Ioc f g).card = (∏ i ∈ f.support ∪ g.support, (Icc (f i) (g i)).card) - 1 := by +theorem card_Ioc : #(Ioc f g) = ∏ i ∈ f.support ∪ g.support, #(Icc (f i) (g i)) - 1 := by rw [card_Ioc_eq_card_Icc_sub_one, card_Icc] -- Porting note: removed [DecidableEq ι] -theorem card_Ioo : (Ioo f g).card = (∏ i ∈ f.support ∪ g.support, (Icc (f i) (g i)).card) - 2 := by +theorem card_Ioo : #(Ioo f g) = ∏ i ∈ f.support ∪ g.support, #(Icc (f i) (g i)) - 2 := by rw [card_Ioo_eq_card_Icc_sub_two, card_Icc] end PartialOrder @@ -119,7 +119,7 @@ variable [Lattice α] [Zero α] [LocallyFiniteOrder α] (f g : ι →₀ α) -- Porting note: removed [DecidableEq ι] theorem card_uIcc : - (uIcc f g).card = ∏ i ∈ f.support ∪ g.support, (uIcc (f i) (g i)).card := by + #(uIcc f g) = ∏ i ∈ f.support ∪ g.support, #(uIcc (f i) (g i)) := by rw [← support_inf_union_support_sup]; exact card_Icc (_ : ι →₀ α) _ end Lattice @@ -129,11 +129,11 @@ section CanonicallyOrdered variable [CanonicallyOrderedAddCommMonoid α] [LocallyFiniteOrder α] variable (f : ι →₀ α) -theorem card_Iic : (Iic f).card = ∏ i ∈ f.support, (Iic (f i)).card := by +theorem card_Iic : #(Iic f) = ∏ i ∈ f.support, #(Iic (f i)) := by classical simp_rw [Iic_eq_Icc, card_Icc, Finsupp.bot_eq_zero, support_zero, empty_union, zero_apply, bot_eq_zero] -theorem card_Iio : (Iio f).card = (∏ i ∈ f.support, (Iic (f i)).card) - 1 := by +theorem card_Iio : #(Iio f) = ∏ i ∈ f.support, #(Iic (f i)) - 1 := by rw [card_Iio_eq_card_Iic_sub_one, card_Iic] end CanonicallyOrdered diff --git a/Mathlib/Data/Fintype/BigOperators.lean b/Mathlib/Data/Fintype/BigOperators.lean index b772a62fd0956..6c8241d4dd465 100644 --- a/Mathlib/Data/Fintype/BigOperators.lean +++ b/Mathlib/Data/Fintype/BigOperators.lean @@ -104,20 +104,20 @@ section Pi variable {ι κ : Type*} {α : ι → Type*} [DecidableEq ι] [DecidableEq κ] @[simp] lemma Finset.card_pi (s : Finset ι) (t : ∀ i, Finset (α i)) : - (s.pi t).card = ∏ i ∈ s, card (t i) := Multiset.card_pi _ _ + #(s.pi t) = ∏ i ∈ s, #(t i) := Multiset.card_pi _ _ namespace Fintype variable [Fintype ι] @[simp] lemma card_piFinset (s : ∀ i, Finset (α i)) : - (piFinset s).card = ∏ i, (s i).card := by simp [piFinset, card_map] + #(piFinset s) = ∏ i, #(s i) := by simp [piFinset, card_map] /-- This lemma is specifically designed to be used backwards, whence the specialisation to `Fin n` as the indexing type doesn't matter in practice. The more general forward direction lemma here is `Fintype.card_piFinset`. -/ lemma card_piFinset_const {α : Type*} (s : Finset α) (n : ℕ) : - (piFinset fun _ : Fin n ↦ s).card = s.card ^ n := by simp + #(piFinset fun _ : Fin n ↦ s) = #s ^ n := by simp @[simp] lemma card_pi [∀ i, Fintype (α i)] : card (∀ i, α i) = ∏ i, card (α i) := card_piFinset _ @@ -132,37 +132,35 @@ lemma card_pi_const (α : Type*) [Fintype α] (n : ℕ) : card (Fin n → α) = card (Sigma α) = ∑ i, card (α i) := card_sigma _ _ /-- The number of dependent maps `f : Π j, s j` for which the `i` component is `a` is the product -over all `j ≠ i` of `(s j).card`. +over all `j ≠ i` of `#(s j)`. Note that this is just a composition of easier lemmas, but there's some glue missing to make that smooth enough not to need this lemma. -/ lemma card_filter_piFinset_eq_of_mem [∀ i, DecidableEq (α i)] (s : ∀ i, Finset (α i)) (i : ι) {a : α i} (ha : a ∈ s i) : - ((piFinset s).filter fun f ↦ f i = a).card = ∏ j ∈ univ.erase i, (s j).card := by + #{f ∈ piFinset s | f i = a} = ∏ j ∈ univ.erase i, #(s j) := by calc - _ = ∏ j, (Function.update s i {a} j).card := by + _ = ∏ j, #(Function.update s i {a} j) := by rw [← piFinset_update_singleton_eq_filter_piFinset_eq _ _ ha, Fintype.card_piFinset] - _ = ∏ j, Function.update (fun j ↦ (s j).card) i 1 j := + _ = ∏ j, Function.update (fun j ↦ #(s j)) i 1 j := Fintype.prod_congr _ _ fun j ↦ by obtain rfl | hji := eq_or_ne j i <;> simp [*] _ = _ := by simp [prod_update_of_mem, erase_eq] lemma card_filter_piFinset_const_eq_of_mem (s : Finset κ) (i : ι) {x : κ} (hx : x ∈ s) : - ((piFinset fun _ ↦ s).filter fun f ↦ f i = x).card = s.card ^ (card ι - 1) := + #{f ∈ piFinset fun _ ↦ s | f i = x} = #s ^ (card ι - 1) := (card_filter_piFinset_eq_of_mem _ _ hx).trans <| by - rw [prod_const s.card, card_erase_of_mem (mem_univ _), card_univ] + rw [prod_const #s, card_erase_of_mem (mem_univ _), card_univ] lemma card_filter_piFinset_eq [∀ i, DecidableEq (α i)] (s : ∀ i, Finset (α i)) (i : ι) (a : α i) : - ((piFinset s).filter fun f ↦ f i = a).card = - if a ∈ s i then ∏ b ∈ univ.erase i, (s b).card else 0 := by + #{f ∈ piFinset s | f i = a} = if a ∈ s i then ∏ b ∈ univ.erase i, #(s b) else 0 := by split_ifs with h · rw [card_filter_piFinset_eq_of_mem _ _ h] · rw [filter_piFinset_of_not_mem _ _ _ h, Finset.card_empty] lemma card_filter_piFinset_const (s : Finset κ) (i : ι) (j : κ) : - ((piFinset fun _ ↦ s).filter fun f ↦ f i = j).card = - if j ∈ s then s.card ^ (card ι - 1) else 0 := + #{f ∈ piFinset fun _ ↦ s | f i = j} = if j ∈ s then #s ^ (card ι - 1) else 0 := (card_filter_piFinset_eq _ _ _).trans <| by - rw [prod_const s.card, card_erase_of_mem (mem_univ _), card_univ] + rw [prod_const #s, card_erase_of_mem (mem_univ _), card_univ] end Fintype end Pi diff --git a/Mathlib/Data/Fintype/Card.lean b/Mathlib/Data/Fintype/Card.lean index 48349c1a8a9b7..1efaf1b99ec29 100644 --- a/Mathlib/Data/Fintype/Card.lean +++ b/Mathlib/Data/Fintype/Card.lean @@ -112,21 +112,21 @@ def truncFinBijection (α) [Fintype α] : Trunc { f : Fin (card α) → α // Bi mem_univ_val univ.2 theorem subtype_card {p : α → Prop} (s : Finset α) (H : ∀ x : α, x ∈ s ↔ p x) : - @card { x // p x } (Fintype.subtype s H) = s.card := + @card { x // p x } (Fintype.subtype s H) = #s := Multiset.card_pmap _ _ _ theorem card_of_subtype {p : α → Prop} (s : Finset α) (H : ∀ x : α, x ∈ s ↔ p x) - [Fintype { x // p x }] : card { x // p x } = s.card := by + [Fintype { x // p x }] : card { x // p x } = #s := by rw [← subtype_card s H] congr! @[simp] theorem card_ofFinset {p : Set α} (s : Finset α) (H : ∀ x, x ∈ s ↔ x ∈ p) : - @Fintype.card p (ofFinset s H) = s.card := + @Fintype.card p (ofFinset s H) = #s := Fintype.subtype_card s H theorem card_of_finset' {p : Set α} (s : Finset α) (H : ∀ x, x ∈ s ↔ x ∈ p) [Fintype p] : - Fintype.card p = s.card := by rw [← card_ofFinset s H]; congr! + Fintype.card p = #s := by rw [← card_ofFinset s H]; congr! end Fintype @@ -223,50 +223,47 @@ theorem toFinset_card {α : Type*} (s : Set α) [Fintype s] : s.toFinset.card = end Set @[simp] -theorem Finset.card_univ [Fintype α] : (Finset.univ : Finset α).card = Fintype.card α := - rfl +theorem Finset.card_univ [Fintype α] : #(univ : Finset α) = Fintype.card α := rfl -theorem Finset.eq_univ_of_card [Fintype α] (s : Finset α) (hs : s.card = Fintype.card α) : +theorem Finset.eq_univ_of_card [Fintype α] (s : Finset α) (hs : #s = Fintype.card α) : s = univ := eq_of_subset_of_card_le (subset_univ _) <| by rw [hs, Finset.card_univ] -theorem Finset.card_eq_iff_eq_univ [Fintype α] (s : Finset α) : - s.card = Fintype.card α ↔ s = Finset.univ := +theorem Finset.card_eq_iff_eq_univ [Fintype α] (s : Finset α) : #s = Fintype.card α ↔ s = univ := ⟨s.eq_univ_of_card, by rintro rfl exact Finset.card_univ⟩ -theorem Finset.card_le_univ [Fintype α] (s : Finset α) : s.card ≤ Fintype.card α := +theorem Finset.card_le_univ [Fintype α] (s : Finset α) : #s ≤ Fintype.card α := card_le_card (subset_univ s) theorem Finset.card_lt_univ_of_not_mem [Fintype α] {s : Finset α} {x : α} (hx : x ∉ s) : - s.card < Fintype.card α := + #s < Fintype.card α := card_lt_card ⟨subset_univ s, not_forall.2 ⟨x, fun hx' => hx (hx' <| mem_univ x)⟩⟩ theorem Finset.card_lt_iff_ne_univ [Fintype α] (s : Finset α) : - s.card < Fintype.card α ↔ s ≠ Finset.univ := + #s < Fintype.card α ↔ s ≠ Finset.univ := s.card_le_univ.lt_iff_ne.trans (not_congr s.card_eq_iff_eq_univ) theorem Finset.card_compl_lt_iff_nonempty [Fintype α] [DecidableEq α] (s : Finset α) : - sᶜ.card < Fintype.card α ↔ s.Nonempty := + #sᶜ < Fintype.card α ↔ s.Nonempty := sᶜ.card_lt_iff_ne_univ.trans s.compl_ne_univ_iff_nonempty theorem Finset.card_univ_diff [DecidableEq α] [Fintype α] (s : Finset α) : - (Finset.univ \ s).card = Fintype.card α - s.card := + #(univ \ s) = Fintype.card α - #s := Finset.card_sdiff (subset_univ s) -theorem Finset.card_compl [DecidableEq α] [Fintype α] (s : Finset α) : - sᶜ.card = Fintype.card α - s.card := +theorem Finset.card_compl [DecidableEq α] [Fintype α] (s : Finset α) : #sᶜ = Fintype.card α - #s := Finset.card_univ_diff s @[simp] theorem Finset.card_add_card_compl [DecidableEq α] [Fintype α] (s : Finset α) : - s.card + sᶜ.card = Fintype.card α := by + #s + #sᶜ = Fintype.card α := by rw [Finset.card_compl, ← Nat.add_sub_assoc (card_le_univ s), Nat.add_sub_cancel_left] @[simp] theorem Finset.card_compl_add_card [DecidableEq α] [Fintype α] (s : Finset α) : - sᶜ.card + s.card = Fintype.card α := by + #sᶜ + #s = Fintype.card α := by rw [add_comm, card_add_card_compl] theorem Fintype.card_compl_set [Fintype α] (s : Set α) [Fintype s] [Fintype (↥sᶜ : Sort _)] : @@ -286,7 +283,7 @@ theorem Fintype.card_fin_lt_of_le {m n : ℕ} (h : m ≤ n) : left_inv := fun i ↦ rfl right_inv := fun i ↦ rfl } -theorem Finset.card_fin (n : ℕ) : Finset.card (Finset.univ : Finset (Fin n)) = n := by simp +theorem Finset.card_fin (n : ℕ) : #(univ : Finset (Fin n)) = n := by simp /-- `Fin` as a map from `ℕ` to `Type` is injective. Note that since this is a statement about equality of types, using it should be avoided if possible. -/ @@ -303,7 +300,7 @@ theorem Fin.cast_eq_cast' {n m : ℕ} (h : Fin n = Fin m) : cases fin_injective h rfl -theorem card_finset_fin_le {n : ℕ} (s : Finset (Fin n)) : s.card ≤ n := by +theorem card_finset_fin_le {n : ℕ} (s : Finset (Fin n)) : #s ≤ n := by simpa only [Fintype.card_fin] using s.card_le_univ theorem Fintype.card_subtype_eq (y : α) [Fintype { x // x = y }] : @@ -529,7 +526,7 @@ theorem card_eq_one_of_forall_eq {i : α} (h : ∀ j, j = i) : card α = 1 := Fintype.card_eq_one_iff.2 ⟨i, h⟩ theorem exists_unique_iff_card_one {α} [Fintype α] (p : α → Prop) [DecidablePred p] : - (∃! a : α, p a) ↔ (Finset.univ.filter p).card = 1 := by + (∃! a : α, p a) ↔ #{x | p x} = 1 := by rw [Finset.card_eq_one] refine exists_congr fun x => ?_ simp only [forall_true_left, Subset.antisymm_iff, subset_singleton_iff', singleton_subset_iff, @@ -648,36 +645,36 @@ def ofRightInverseOfCardLE (hαβ : card α ≤ card β) (f : α → β) (g : β end Equiv @[simp] -theorem Fintype.card_coe (s : Finset α) [Fintype s] : Fintype.card s = s.card := +theorem Fintype.card_coe (s : Finset α) [Fintype s] : Fintype.card s = #s := @Fintype.card_of_finset' _ _ _ (fun _ => Iff.rfl) (id _) -/-- Noncomputable equivalence between a finset `s` coerced to a type and `Fin s.card`. -/ -noncomputable def Finset.equivFin (s : Finset α) : s ≃ Fin s.card := +/-- Noncomputable equivalence between a finset `s` coerced to a type and `Fin #s`. -/ +noncomputable def Finset.equivFin (s : Finset α) : s ≃ Fin #s := Fintype.equivFinOfCardEq (Fintype.card_coe _) /-- Noncomputable equivalence between a finset `s` as a fintype and `Fin n`, when there is a -proof that `s.card = n`. -/ -noncomputable def Finset.equivFinOfCardEq {s : Finset α} {n : ℕ} (h : s.card = n) : s ≃ Fin n := +proof that `#s = n`. -/ +noncomputable def Finset.equivFinOfCardEq {s : Finset α} {n : ℕ} (h : #s = n) : s ≃ Fin n := Fintype.equivFinOfCardEq ((Fintype.card_coe _).trans h) -theorem Finset.card_eq_of_equiv_fin {s : Finset α} {n : ℕ} (i : s ≃ Fin n) : s.card = n := +theorem Finset.card_eq_of_equiv_fin {s : Finset α} {n : ℕ} (i : s ≃ Fin n) : #s = n := Fin.equiv_iff_eq.1 ⟨s.equivFin.symm.trans i⟩ theorem Finset.card_eq_of_equiv_fintype {s : Finset α} [Fintype β] (i : s ≃ β) : - s.card = Fintype.card β := card_eq_of_equiv_fin <| i.trans <| Fintype.equivFin β + #s = Fintype.card β := card_eq_of_equiv_fin <| i.trans <| Fintype.equivFin β /-- Noncomputable equivalence between two finsets `s` and `t` as fintypes when there is a proof -that `s.card = t.card`. -/ -noncomputable def Finset.equivOfCardEq {s : Finset α} {t : Finset β} (h : s.card = t.card) : +that `#s = #t`. -/ +noncomputable def Finset.equivOfCardEq {s : Finset α} {t : Finset β} (h : #s = #t) : s ≃ t := Fintype.equivOfCardEq ((Fintype.card_coe _).trans (h.trans (Fintype.card_coe _).symm)) -theorem Finset.card_eq_of_equiv {s : Finset α} {t : Finset β} (i : s ≃ t) : s.card = t.card := +theorem Finset.card_eq_of_equiv {s : Finset α} {t : Finset β} (i : s ≃ t) : #s = #t := (card_eq_of_equiv_fintype i).trans (Fintype.card_coe _) /-- We can inflate a set `s` to any bigger size. -/ -lemma Finset.exists_superset_card_eq [Fintype α] {n : ℕ} {s : Finset α} (hsn : s.card ≤ n) +lemma Finset.exists_superset_card_eq [Fintype α] {n : ℕ} {s : Finset α} (hsn : #s ≤ n) (hnα : n ≤ Fintype.card α) : - ∃ t, s ⊆ t ∧ t.card = n := by simpa using exists_subsuperset_card_eq s.subset_univ hsn hnα + ∃ t, s ⊆ t ∧ #t = n := by simpa using exists_subsuperset_card_eq s.subset_univ hsn hnα @[simp] theorem Fintype.card_prop : Fintype.card Prop = 2 := @@ -728,7 +725,7 @@ theorem nonempty_iff_card_le [Fintype α] [Fintype β] : Nonempty (α ↪ β) ↔ Fintype.card α ≤ Fintype.card β := ⟨fun ⟨e⟩ => Fintype.card_le_of_embedding e, nonempty_of_card_le⟩ -theorem exists_of_card_le_finset [Fintype α] {s : Finset β} (h : Fintype.card α ≤ s.card) : +theorem exists_of_card_le_finset [Fintype α] {s : Finset β} (h : Fintype.card α ≤ #s) : ∃ f : α ↪ β, Set.range f ⊆ s := by rw [← Fintype.card_coe] at h rcases nonempty_of_card_le h with ⟨f⟩ @@ -760,7 +757,7 @@ theorem Fintype.card_subtype_lt [Fintype α] {p : α → Prop} [DecidablePred p] rwa [Subtype.range_coe_subtype] theorem Fintype.card_subtype [Fintype α] (p : α → Prop) [DecidablePred p] : - Fintype.card { x // p x } = ((Finset.univ : Finset α).filter p).card := by + Fintype.card { x // p x } = #{x | p x} := by refine Fintype.card_of_subtype _ ?_ simp @@ -810,9 +807,7 @@ theorem wellFounded_of_trans_of_irrefl (r : α → α → Prop) [IsTrans α r] [ WellFounded r := by classical cases nonempty_fintype α - have : - ∀ x y, r x y → (univ.filter fun z => r z x).card < (univ.filter fun z => r z y).card := - fun x y hxy => + have (x y) (hxy : r x y) : #{z | r z x} < #{z | r z y} := Finset.card_lt_card <| by simp only [Finset.lt_iff_ssubset.symm, lt_iff_le_not_le, Finset.le_iff_subset, Finset.subset_iff, mem_filter, true_and, mem_univ, hxy] @@ -884,7 +879,7 @@ theorem of_injective_to_set {s : Set α} (hs : s ≠ Set.univ) {f : α → s} (h refine lt_irrefl (Fintype.card α) ?_ calc Fintype.card α ≤ Fintype.card s := Fintype.card_le_of_injective f hf - _ = s.toFinset.card := s.toFinset_card.symm + _ = #s.toFinset := s.toFinset_card.symm _ < Fintype.card α := Finset.card_lt_card <| by rwa [Set.toFinset_ssubset_univ, Set.ssubset_univ_iff] @@ -997,14 +992,14 @@ noncomputable def natEmbedding (α : Type*) [Infinite α] : ℕ ↪ α := ⟨_, natEmbeddingAux_injective α⟩ /-- See `Infinite.exists_superset_card_eq` for a version that, for an `s : Finset α`, -provides a superset `t : Finset α`, `s ⊆ t` such that `t.card` is fixed. -/ -theorem exists_subset_card_eq (α : Type*) [Infinite α] (n : ℕ) : ∃ s : Finset α, s.card = n := +provides a superset `t : Finset α`, `s ⊆ t` such that `#t` is fixed. -/ +theorem exists_subset_card_eq (α : Type*) [Infinite α] (n : ℕ) : ∃ s : Finset α, #s = n := ⟨(range n).map (natEmbedding α), by rw [card_map, card_range]⟩ /-- See `Infinite.exists_subset_card_eq` for a version that provides an arbitrary `s : Finset α` for any cardinality. -/ -theorem exists_superset_card_eq [Infinite α] (s : Finset α) (n : ℕ) (hn : s.card ≤ n) : - ∃ t : Finset α, s ⊆ t ∧ t.card = n := by +theorem exists_superset_card_eq [Infinite α] (s : Finset α) (n : ℕ) (hn : #s ≤ n) : + ∃ t : Finset α, s ⊆ t ∧ #t = n := by induction' n with n IH generalizing s · exact ⟨s, subset_refl _, Nat.eq_zero_of_le_zero hn⟩ · rcases hn.eq_or_lt with hn' | hn' @@ -1017,7 +1012,7 @@ theorem exists_superset_card_eq [Infinite α] (s : Finset α) (n : ℕ) (hn : s. end Infinite /-- If every finset in a type has bounded cardinality, that type is finite. -/ -noncomputable def fintypeOfFinsetCardLe {ι : Type*} (n : ℕ) (w : ∀ s : Finset ι, s.card ≤ n) : +noncomputable def fintypeOfFinsetCardLe {ι : Type*} (n : ℕ) (w : ∀ s : Finset ι, #s ≤ n) : Fintype ι := by apply fintypeOfNotInfinite intro i diff --git a/Mathlib/Data/Fintype/Fin.lean b/Mathlib/Data/Fintype/Fin.lean index 9c66b04ae7c46..63559a8ab0c85 100644 --- a/Mathlib/Data/Fintype/Fin.lean +++ b/Mathlib/Data/Fintype/Fin.lean @@ -53,19 +53,18 @@ theorem Iio_castSucc (i : Fin n) : Iio (castSucc i) = (Iio i).map Fin.castSuccEm exact (Fin.map_valEmbedding_Iio i).symm theorem card_filter_univ_succ (p : Fin (n + 1) → Prop) [DecidablePred p] : - (univ.filter p).card = - if p 0 then (univ.filter (p ∘ Fin.succ)).card + 1 else (univ.filter (p ∘ Fin.succ)).card := by + #{x | p x} = if p 0 then #{x | p (.succ x)} + 1 else #{x | p (.succ x)} := by rw [Fin.univ_succ, filter_cons, apply_ite Finset.card, card_cons, filter_map, card_map]; rfl theorem card_filter_univ_succ' (p : Fin (n + 1) → Prop) [DecidablePred p] : - (univ.filter p).card = ite (p 0) 1 0 + (univ.filter (p ∘ Fin.succ)).card := by + #{x | p x} = ite (p 0) 1 0 + #{x | p (.succ x)}:= by rw [card_filter_univ_succ]; split_ifs <;> simp [add_comm] theorem card_filter_univ_eq_vector_get_eq_count [DecidableEq α] (a : α) (v : Vector α n) : - (univ.filter fun i => v.get i = a).card = v.toList.count a := by + #{i | v.get i = a} = v.toList.count a := by induction' v with n x xs hxs · simp - · simp_rw [card_filter_univ_succ', Vector.get_cons_zero, Vector.toList_cons, Function.comp_def, - Vector.get_cons_succ, hxs, List.count_cons, add_comm (ite (x = a) 1 0), beq_iff_eq] + · simp_rw [card_filter_univ_succ', Vector.get_cons_zero, Vector.toList_cons, Vector.get_cons_succ, + hxs, List.count_cons, add_comm (ite (x = a) 1 0), beq_iff_eq] end Fin diff --git a/Mathlib/Data/Fintype/Perm.lean b/Mathlib/Data/Fintype/Perm.lean index a37481050a9ef..823a08e377bc8 100644 --- a/Mathlib/Data/Fintype/Perm.lean +++ b/Mathlib/Data/Fintype/Perm.lean @@ -133,7 +133,7 @@ theorem mem_perms_of_finset_iff : ∀ {s : Finset α} {f : Perm α}, f ∈ permsOfFinset s ↔ ∀ {x}, f x ≠ x → x ∈ s := by rintro ⟨⟨l⟩, hs⟩ f; exact mem_permsOfList_iff -theorem card_perms_of_finset : ∀ s : Finset α, (permsOfFinset s).card = s.card ! := by +theorem card_perms_of_finset : ∀ s : Finset α, #(permsOfFinset s) = (#s)! := by rintro ⟨⟨l⟩, hs⟩; exact length_permsOfList l /-- The collection of permutations of a fintype is a fintype. -/ diff --git a/Mathlib/Data/Fintype/Pi.lean b/Mathlib/Data/Fintype/Pi.lean index c32f9bbfd62b6..8fdb303b7f64b 100644 --- a/Mathlib/Data/Fintype/Pi.lean +++ b/Mathlib/Data/Fintype/Pi.lean @@ -104,12 +104,12 @@ lemma eval_image_piFinset_const {β} [DecidableEq β] (t : Finset β) (a : α) : variable [∀ a, DecidableEq (δ a)] lemma filter_piFinset_of_not_mem (t : ∀ a, Finset (δ a)) (a : α) (x : δ a) (hx : x ∉ t a) : - (piFinset t).filter (· a = x) = ∅ := by + {f ∈ piFinset t | f a = x} = ∅ := by simp only [filter_eq_empty_iff, mem_piFinset]; rintro f hf rfl; exact hx (hf _) -- TODO: This proof looks like a good example of something that `aesop` can't do but should lemma piFinset_update_eq_filter_piFinset_mem (s : ∀ i, Finset (δ i)) (i : α) {t : Finset (δ i)} - (hts : t ⊆ s i) : piFinset (Function.update s i t) = (piFinset s).filter (fun f ↦ f i ∈ t) := by + (hts : t ⊆ s i) : piFinset (Function.update s i t) = {f ∈ piFinset s | f i ∈ t} := by ext f simp only [mem_piFinset, mem_filter] refine ⟨fun h ↦ ?_, fun h j ↦ ?_⟩ @@ -124,7 +124,7 @@ lemma piFinset_update_eq_filter_piFinset_mem (s : ∀ i, Finset (δ i)) (i : α) lemma piFinset_update_singleton_eq_filter_piFinset_eq (s : ∀ i, Finset (δ i)) (i : α) {a : δ i} (ha : a ∈ s i) : - piFinset (Function.update s i {a}) = (piFinset s).filter (fun f ↦ f i = a) := by + piFinset (Function.update s i {a}) = {f ∈ piFinset s | f i = a} := by simp [piFinset_update_eq_filter_piFinset_mem, ha] end Fintype @@ -172,7 +172,7 @@ namespace Finset variable {ι : Type*} [DecidableEq (ι → α)] {s : Finset α} {f : ι → α} lemma piFinset_filter_const [DecidableEq ι] [Fintype ι] : - (Fintype.piFinset fun _ ↦ s).filter (∃ a ∈ s, const ι a = ·) = s.piDiag ι := by aesop + {f ∈ Fintype.piFinset fun _ : ι ↦ s | ∃ a ∈ s, const ι a = f} = s.piDiag ι := by aesop lemma piDiag_subset_piFinset [DecidableEq ι] [Fintype ι] : s.piDiag ι ⊆ Fintype.piFinset fun _ ↦ s := by simp [← piFinset_filter_const] diff --git a/Mathlib/Data/Fintype/Powerset.lean b/Mathlib/Data/Fintype/Powerset.lean index 206d1da75abc3..7e8604fc99236 100644 --- a/Mathlib/Data/Fintype/Powerset.lean +++ b/Mathlib/Data/Fintype/Powerset.lean @@ -29,30 +29,30 @@ variable [Fintype α] {s : Finset α} {k : ℕ} coe_injective <| by simp [-coe_eq_univ] lemma filter_subset_univ [DecidableEq α] (s : Finset α) : - filter (fun t ↦ t ⊆ s) univ = powerset s := by ext; simp + ({t | t ⊆ s} : Finset _) = powerset s := by ext; simp @[simp] lemma powerset_eq_univ : s.powerset = univ ↔ s = univ := by rw [← Finset.powerset_univ, powerset_inj] -lemma mem_powersetCard_univ : s ∈ powersetCard k (univ : Finset α) ↔ card s = k := +lemma mem_powersetCard_univ : s ∈ powersetCard k (univ : Finset α) ↔ #s = k := mem_powersetCard.trans <| and_iff_right <| subset_univ _ variable (α) @[simp] lemma univ_filter_card_eq (k : ℕ) : - (univ : Finset (Finset α)).filter (fun s ↦ s.card = k) = univ.powersetCard k := by ext; simp + ({s | #s = k} : Finset (Finset α)) = univ.powersetCard k := by ext; simp end Finset @[simp] theorem Fintype.card_finset_len [Fintype α] (k : ℕ) : - Fintype.card { s : Finset α // s.card = k } = Nat.choose (Fintype.card α) k := by + Fintype.card { s : Finset α // #s = k } = Nat.choose (Fintype.card α) k := by simp [Fintype.subtype_card, Finset.card_univ] instance Set.fintype [Fintype α] : Fintype (Set α) := ⟨(@Finset.univ (Finset α) _).map coeEmb.1, fun s => by classical - refine mem_map.2 ⟨Finset.univ.filter (· ∈ s), Finset.mem_univ _, (coe_filter _ _).trans ?_⟩ + refine mem_map.2 ⟨({a | a ∈ s} : Finset _), Finset.mem_univ _, (coe_filter _ _).trans ?_⟩ simp⟩ -- Not to be confused with `Set.Finite`, the predicate diff --git a/Mathlib/Data/Fintype/Sort.lean b/Mathlib/Data/Fintype/Sort.lean index 40f2ba4db88c3..de5153b97b265 100644 --- a/Mathlib/Data/Fintype/Sort.lean +++ b/Mathlib/Data/Fintype/Sort.lean @@ -32,7 +32,7 @@ variable {α : Type*} [DecidableEq α] [Fintype α] [LinearOrder α] {m n : ℕ} cardinality `n`, then `Fin m ⊕ Fin n ≃ α`. The equivalence sends elements of `Fin m` to elements of `s` and elements of `Fin n` to elements of `sᶜ` while preserving order on each "half" of `Fin m ⊕ Fin n` (using `Set.orderIsoOfFin`). -/ -def finSumEquivOfFinset (hm : s.card = m) (hn : sᶜ.card = n) : Fin m ⊕ Fin n ≃ α := +def finSumEquivOfFinset (hm : #s = m) (hn : #sᶜ = n) : Fin m ⊕ Fin n ≃ α := calc Fin m ⊕ Fin n ≃ (s : Set α) ⊕ (sᶜ : Set α) := Equiv.sumCongr (s.orderIsoOfFin hm).toEquiv <| @@ -40,11 +40,11 @@ def finSumEquivOfFinset (hm : s.card = m) (hn : sᶜ.card = n) : Fin m ⊕ Fin n _ ≃ α := Equiv.Set.sumCompl _ @[simp] -theorem finSumEquivOfFinset_inl (hm : s.card = m) (hn : sᶜ.card = n) (i : Fin m) : +theorem finSumEquivOfFinset_inl (hm : #s = m) (hn : #sᶜ = n) (i : Fin m) : finSumEquivOfFinset hm hn (Sum.inl i) = s.orderEmbOfFin hm i := rfl @[simp] -theorem finSumEquivOfFinset_inr (hm : s.card = m) (hn : sᶜ.card = n) (i : Fin n) : +theorem finSumEquivOfFinset_inr (hm : #s = m) (hn : #sᶜ = n) (i : Fin n) : finSumEquivOfFinset hm hn (Sum.inr i) = sᶜ.orderEmbOfFin hn i := rfl diff --git a/Mathlib/Data/Fintype/Sum.lean b/Mathlib/Data/Fintype/Sum.lean index 371313c49a376..930f4d0abf5d9 100644 --- a/Mathlib/Data/Fintype/Sum.lean +++ b/Mathlib/Data/Fintype/Sum.lean @@ -70,7 +70,7 @@ theorem image_subtype_univ_ssubset_image_univ [Fintype α] [DecidableEq β] (k : /-- Any injection from a finset `s` in a fintype `α` to a finset `t` of the same cardinality as `α` can be extended to a bijection between `α` and `t`. -/ theorem Finset.exists_equiv_extend_of_card_eq [Fintype α] [DecidableEq β] {t : Finset β} - (hαt : Fintype.card α = t.card) {s : Finset α} {f : α → β} (hfst : Finset.image f s ⊆ t) + (hαt : Fintype.card α = #t) {s : Finset α} {f : α → β} (hfst : Finset.image f s ⊆ t) (hfs : Set.InjOn f s) : ∃ g : α ≃ t, ∀ i ∈ s, (g i : β) = f i := by classical induction' s using Finset.induction with a s has H generalizing f @@ -95,7 +95,7 @@ theorem Finset.exists_equiv_extend_of_card_eq [Fintype α] [DecidableEq β] {t : /-- Any injection from a set `s` in a fintype `α` to a finset `t` of the same cardinality as `α` can be extended to a bijection between `α` and `t`. -/ theorem Set.MapsTo.exists_equiv_extend_of_card_eq [Fintype α] {t : Finset β} - (hαt : Fintype.card α = t.card) {s : Set α} {f : α → β} (hfst : s.MapsTo f t) + (hαt : Fintype.card α = #t) {s : Set α} {f : α → β} (hfst : s.MapsTo f t) (hfs : Set.InjOn f s) : ∃ g : α ≃ t, ∀ i ∈ s, (g i : β) = f i := by classical let s' : Finset α := s.toFinset diff --git a/Mathlib/Data/Int/CardIntervalMod.lean b/Mathlib/Data/Int/CardIntervalMod.lean index b5725e3f43eb6..b243bdf2208d0 100644 --- a/Mathlib/Data/Int/CardIntervalMod.lean +++ b/Mathlib/Data/Int/CardIntervalMod.lean @@ -25,14 +25,16 @@ namespace Int variable (a b : ℤ) {r : ℤ} -lemma Ico_filter_modEq_eq (v : ℤ) : (Ico a b).filter (· ≡ v [ZMOD r]) = - ((Ico (a - v) (b - v)).filter (r ∣ ·)).map ⟨(· + v), add_left_injective v⟩ := by +lemma Ico_filter_modEq_eq (v : ℤ) : + {x ∈ Ico a b | x ≡ v [ZMOD r]} = + {x ∈ Ico (a - v) (b - v) | r ∣ x}.map ⟨(· + v), add_left_injective v⟩ := by ext x simp_rw [mem_map, mem_filter, mem_Ico, Function.Embedding.coeFn_mk, ← eq_sub_iff_add_eq, exists_eq_right, modEq_comm, modEq_iff_dvd, sub_lt_sub_iff_right, sub_le_sub_iff_right] -lemma Ioc_filter_modEq_eq (v : ℤ) : (Ioc a b).filter (· ≡ v [ZMOD r]) = - ((Ioc (a - v) (b - v)).filter (r ∣ ·)).map ⟨(· + v), add_left_injective v⟩ := by +lemma Ioc_filter_modEq_eq (v : ℤ) : + {x ∈ Ioc a b | x ≡ v [ZMOD r]} = + {x ∈ Ioc (a - v) (b - v) | r ∣ x}.map ⟨(· + v), add_left_injective v⟩ := by ext x simp_rw [mem_map, mem_filter, mem_Ioc, Function.Embedding.coeFn_mk, ← eq_sub_iff_add_eq, exists_eq_right, modEq_comm, modEq_iff_dvd, sub_lt_sub_iff_right, sub_le_sub_iff_right] @@ -40,40 +42,40 @@ lemma Ioc_filter_modEq_eq (v : ℤ) : (Ioc a b).filter (· ≡ v [ZMOD r]) = variable (hr : 0 < r) include hr -lemma Ico_filter_dvd_eq : (Ico a b).filter (r ∣ ·) = - (Ico ⌈a / (r : ℚ)⌉ ⌈b / (r : ℚ)⌉).map ⟨(· * r), mul_left_injective₀ hr.ne'⟩ := by +lemma Ico_filter_dvd_eq : + {x ∈ Ico a b | r ∣ x} = + (Ico ⌈a / (r : ℚ)⌉ ⌈b / (r : ℚ)⌉).map ⟨(· * r), mul_left_injective₀ hr.ne'⟩ := by ext x simp only [mem_map, mem_filter, mem_Ico, ceil_le, lt_ceil, div_le_iff₀, lt_div_iff₀, dvd_iff_exists_eq_mul_left, cast_pos.2 hr, ← cast_mul, cast_lt, cast_le] aesop -lemma Ioc_filter_dvd_eq : (Ioc a b).filter (r ∣ ·) = - (Ioc ⌊a / (r : ℚ)⌋ ⌊b / (r : ℚ)⌋).map ⟨(· * r), mul_left_injective₀ hr.ne'⟩ := by +lemma Ioc_filter_dvd_eq : + {x ∈ Ioc a b | r ∣ x} = + (Ioc ⌊a / (r : ℚ)⌋ ⌊b / (r : ℚ)⌋).map ⟨(· * r), mul_left_injective₀ hr.ne'⟩ := by ext x simp only [mem_map, mem_filter, mem_Ioc, floor_lt, le_floor, div_lt_iff₀, le_div_iff₀, dvd_iff_exists_eq_mul_left, cast_pos.2 hr, ← cast_mul, cast_lt, cast_le] aesop /-- There are `⌈b / r⌉ - ⌈a / r⌉` multiples of `r` in `[a, b)`, if `a ≤ b`. -/ -theorem Ico_filter_dvd_card : ((Ico a b).filter (r ∣ ·)).card = - max (⌈b / (r : ℚ)⌉ - ⌈a / (r : ℚ)⌉) 0 := by +theorem Ico_filter_dvd_card : #{x ∈ Ico a b | r ∣ x} = max (⌈b / (r : ℚ)⌉ - ⌈a / (r : ℚ)⌉) 0 := by rw [Ico_filter_dvd_eq _ _ hr, card_map, card_Ico, toNat_eq_max] /-- There are `⌊b / r⌋ - ⌊a / r⌋` multiples of `r` in `(a, b]`, if `a ≤ b`. -/ -theorem Ioc_filter_dvd_card : ((Ioc a b).filter (r ∣ ·)).card = - max (⌊b / (r : ℚ)⌋ - ⌊a / (r : ℚ)⌋) 0 := by +theorem Ioc_filter_dvd_card : #{x ∈ Ioc a b | r ∣ x} = max (⌊b / (r : ℚ)⌋ - ⌊a / (r : ℚ)⌋) 0 := by rw [Ioc_filter_dvd_eq _ _ hr, card_map, card_Ioc, toNat_eq_max] /-- There are `⌈(b - v) / r⌉ - ⌈(a - v) / r⌉` numbers congruent to `v` mod `r` in `[a, b)`, if `a ≤ b`. -/ -theorem Ico_filter_modEq_card (v : ℤ) : ((Ico a b).filter (· ≡ v [ZMOD r])).card = - max (⌈(b - v) / (r : ℚ)⌉ - ⌈(a - v) / (r : ℚ)⌉) 0 := by +theorem Ico_filter_modEq_card (v : ℤ) : + #{x ∈ Ico a b | x ≡ v [ZMOD r]} = max (⌈(b - v) / (r : ℚ)⌉ - ⌈(a - v) / (r : ℚ)⌉) 0 := by simp [Ico_filter_modEq_eq, Ico_filter_dvd_eq, toNat_eq_max, hr] /-- There are `⌊(b - v) / r⌋ - ⌊(a - v) / r⌋` numbers congruent to `v` mod `r` in `(a, b]`, if `a ≤ b`. -/ -theorem Ioc_filter_modEq_card (v : ℤ) : ((Ioc a b).filter (· ≡ v [ZMOD r])).card = - max (⌊(b - v) / (r : ℚ)⌋ - ⌊(a - v) / (r : ℚ)⌋) 0 := by +theorem Ioc_filter_modEq_card (v : ℤ) : + #{x ∈ Ioc a b | x ≡ v [ZMOD r]} = max (⌊(b - v) / (r : ℚ)⌋ - ⌊(a - v) / (r : ℚ)⌋) 0 := by simp [Ioc_filter_modEq_eq, Ioc_filter_dvd_eq, toNat_eq_max, hr] end Int @@ -82,16 +84,18 @@ namespace Nat variable (a b : ℕ) {r : ℕ} -lemma Ico_filter_modEq_cast {v : ℕ} : ((Ico a b).filter (· ≡ v [MOD r])).map castEmbedding = - (Ico (a : ℤ) (b : ℤ)).filter (· ≡ v [ZMOD r]) := by +lemma Ico_filter_modEq_cast {v : ℕ} : + {x ∈ Ico a b | x ≡ v [MOD r]}.map castEmbedding = + {x ∈ Ico (a : ℤ) (b : ℤ) | x ≡ v [ZMOD r]} := by ext x simp only [mem_map, mem_filter, mem_Ico, castEmbedding_apply] constructor · simp_rw [forall_exists_index, ← natCast_modEq_iff]; intro y ⟨h, c⟩; subst c; exact_mod_cast h · intro h; lift x to ℕ using (by linarith); exact ⟨x, by simp_all [natCast_modEq_iff]⟩ -lemma Ioc_filter_modEq_cast {v : ℕ} : ((Ioc a b).filter (· ≡ v [MOD r])).map castEmbedding = - (Ioc (a : ℤ) (b : ℤ)).filter (· ≡ v [ZMOD r]) := by +lemma Ioc_filter_modEq_cast {v : ℕ} : + {x ∈ Ioc a b | x ≡ v [MOD r]}.map castEmbedding = + {x ∈ Ioc (a : ℤ) (b : ℤ) | x ≡ v [ZMOD r]} := by ext x simp only [mem_map, mem_filter, mem_Ioc, castEmbedding_apply] constructor @@ -103,15 +107,15 @@ include hr /-- There are `⌈(b - v) / r⌉ - ⌈(a - v) / r⌉` numbers congruent to `v` mod `r` in `[a, b)`, if `a ≤ b`. `Nat` version of `Int.Ico_filter_modEq_card`. -/ -theorem Ico_filter_modEq_card (v : ℕ) : ((Ico a b).filter (· ≡ v [MOD r])).card = - max (⌈(b - v) / (r : ℚ)⌉ - ⌈(a - v) / (r : ℚ)⌉) 0 := by +theorem Ico_filter_modEq_card (v : ℕ) : + #{x ∈ Ico a b | x ≡ v [MOD r]} = max (⌈(b - v) / (r : ℚ)⌉ - ⌈(a - v) / (r : ℚ)⌉) 0 := by simp_rw [← Ico_filter_modEq_cast _ _ ▸ card_map _, Int.Ico_filter_modEq_card _ _ (cast_lt.mpr hr), Int.cast_natCast] /-- There are `⌊(b - v) / r⌋ - ⌊(a - v) / r⌋` numbers congruent to `v` mod `r` in `(a, b]`, if `a ≤ b`. `Nat` version of `Int.Ioc_filter_modEq_card`. -/ -theorem Ioc_filter_modEq_card (v : ℕ) : ((Ioc a b).filter (· ≡ v [MOD r])).card = - max (⌊(b - v) / (r : ℚ)⌋ - ⌊(a - v) / (r : ℚ)⌋) 0 := by +theorem Ioc_filter_modEq_card (v : ℕ) : + #{x ∈ Ioc a b | x ≡ v [MOD r]} = max (⌊(b - v) / (r : ℚ)⌋ - ⌊(a - v) / (r : ℚ)⌋) 0 := by simp_rw [← Ioc_filter_modEq_cast _ _ ▸ card_map _, Int.Ioc_filter_modEq_card _ _ (cast_lt.mpr hr), Int.cast_natCast] diff --git a/Mathlib/Data/Int/Interval.lean b/Mathlib/Data/Int/Interval.lean index ce53928746469..afb891ee30e61 100644 --- a/Mathlib/Data/Int/Interval.lean +++ b/Mathlib/Data/Int/Interval.lean @@ -98,35 +98,35 @@ theorem uIcc_eq_finset_map : (Nat.castEmbedding.trans <| addLeftEmbedding <| min a b) := rfl @[simp] -theorem card_Icc : (Icc a b).card = (b + 1 - a).toNat := (card_map _).trans <| card_range _ +theorem card_Icc : #(Icc a b) = (b + 1 - a).toNat := (card_map _).trans <| card_range _ @[simp] -theorem card_Ico : (Ico a b).card = (b - a).toNat := (card_map _).trans <| card_range _ +theorem card_Ico : #(Ico a b) = (b - a).toNat := (card_map _).trans <| card_range _ @[simp] -theorem card_Ioc : (Ioc a b).card = (b - a).toNat := (card_map _).trans <| card_range _ +theorem card_Ioc : #(Ioc a b) = (b - a).toNat := (card_map _).trans <| card_range _ @[simp] -theorem card_Ioo : (Ioo a b).card = (b - a - 1).toNat := (card_map _).trans <| card_range _ +theorem card_Ioo : #(Ioo a b) = (b - a - 1).toNat := (card_map _).trans <| card_range _ @[simp] -theorem card_uIcc : (uIcc a b).card = (b - a).natAbs + 1 := +theorem card_uIcc : #(uIcc a b) = (b - a).natAbs + 1 := (card_map _).trans <| (Nat.cast_inj (R := ℤ)).mp <| by rw [card_range, sup_eq_max, inf_eq_min, Int.toNat_of_nonneg (sub_nonneg_of_le <| le_add_one min_le_max), Int.ofNat_add, Int.natCast_natAbs, add_comm, add_sub_assoc, max_sub_min_eq_abs, add_comm, Int.ofNat_one] -theorem card_Icc_of_le (h : a ≤ b + 1) : ((Icc a b).card : ℤ) = b + 1 - a := by +theorem card_Icc_of_le (h : a ≤ b + 1) : (#(Icc a b) : ℤ) = b + 1 - a := by rw [card_Icc, toNat_sub_of_le h] -theorem card_Ico_of_le (h : a ≤ b) : ((Ico a b).card : ℤ) = b - a := by +theorem card_Ico_of_le (h : a ≤ b) : (#(Ico a b) : ℤ) = b - a := by rw [card_Ico, toNat_sub_of_le h] -theorem card_Ioc_of_le (h : a ≤ b) : ((Ioc a b).card : ℤ) = b - a := by +theorem card_Ioc_of_le (h : a ≤ b) : (#(Ioc a b) : ℤ) = b - a := by rw [card_Ioc, toNat_sub_of_le h] -theorem card_Ioo_of_lt (h : a < b) : ((Ioo a b).card : ℤ) = b - a - 1 := by +theorem card_Ioo_of_lt (h : a < b) : (#(Ioo a b) : ℤ) = b - a - 1 := by rw [card_Ioo, sub_sub, toNat_sub_of_le h] -- Porting note (#11119): removed `simp` attribute because `simpNF` says it can prove it diff --git a/Mathlib/Data/Multiset/Interval.lean b/Mathlib/Data/Multiset/Interval.lean index d4db3e807d935..6602425bbde82 100644 --- a/Mathlib/Data/Multiset/Interval.lean +++ b/Mathlib/Data/Multiset/Interval.lean @@ -51,20 +51,20 @@ theorem uIcc_eq : (Icc_eq _ _).trans <| by simp [uIcc] theorem card_Icc : - (Finset.Icc s t).card = ∏ i ∈ s.toFinset ∪ t.toFinset, (t.count i + 1 - s.count i) := by + #(Finset.Icc s t) = ∏ i ∈ s.toFinset ∪ t.toFinset, (t.count i + 1 - s.count i) := by simp_rw [Icc_eq, Finset.card_map, DFinsupp.card_Icc, Nat.card_Icc, Multiset.toDFinsupp_apply, toDFinsupp_support] theorem card_Ico : - (Finset.Ico s t).card = ∏ i ∈ s.toFinset ∪ t.toFinset, (t.count i + 1 - s.count i) - 1 := by + #(Finset.Ico s t) = ∏ i ∈ s.toFinset ∪ t.toFinset, (t.count i + 1 - s.count i) - 1 := by rw [Finset.card_Ico_eq_card_Icc_sub_one, card_Icc] theorem card_Ioc : - (Finset.Ioc s t).card = ∏ i ∈ s.toFinset ∪ t.toFinset, (t.count i + 1 - s.count i) - 1 := by + #(Finset.Ioc s t) = ∏ i ∈ s.toFinset ∪ t.toFinset, (t.count i + 1 - s.count i) - 1 := by rw [Finset.card_Ioc_eq_card_Icc_sub_one, card_Icc] theorem card_Ioo : - (Finset.Ioo s t).card = ∏ i ∈ s.toFinset ∪ t.toFinset, (t.count i + 1 - s.count i) - 2 := by + #(Finset.Ioo s t) = ∏ i ∈ s.toFinset ∪ t.toFinset, (t.count i + 1 - s.count i) - 2 := by rw [Finset.card_Ioo_eq_card_Icc_sub_two, card_Icc] theorem card_uIcc : diff --git a/Mathlib/Data/Nat/Choose/Sum.lean b/Mathlib/Data/Nat/Choose/Sum.lean index ab6812bc6fcb7..fd5c6f3f73caa 100644 --- a/Mathlib/Data/Nat/Choose/Sum.lean +++ b/Mathlib/Data/Nat/Choose/Sum.lean @@ -168,8 +168,8 @@ theorem Int.alternating_sum_range_choose_of_ne {n : ℕ} (h0 : n ≠ 0) : namespace Finset theorem sum_powerset_apply_card {α β : Type*} [AddCommMonoid α] (f : ℕ → α) {x : Finset β} : - ∑ m ∈ x.powerset, f m.card = ∑ m ∈ range (x.card + 1), x.card.choose m • f m := by - trans ∑ m ∈ range (x.card + 1), ∑ j ∈ x.powerset.filter fun z ↦ z.card = m, f j.card + ∑ m ∈ x.powerset, f #m = ∑ m ∈ range (#x + 1), (#x).choose m • f m := by + trans ∑ m ∈ range (#x + 1), ∑ j ∈ x.powerset with #j = m, f #j · refine (sum_fiberwise_of_maps_to ?_ _).symm intro y hy rw [mem_range, Nat.lt_succ_iff] @@ -181,12 +181,12 @@ theorem sum_powerset_apply_card {α β : Type*} [AddCommMonoid α] (f : ℕ → rw [(mem_powersetCard.1 hz).2] theorem sum_powerset_neg_one_pow_card {α : Type*} [DecidableEq α] {x : Finset α} : - (∑ m ∈ x.powerset, (-1 : ℤ) ^ m.card) = if x = ∅ then 1 else 0 := by + (∑ m ∈ x.powerset, (-1 : ℤ) ^ #m) = if x = ∅ then 1 else 0 := by rw [sum_powerset_apply_card] simp only [nsmul_eq_mul', ← card_eq_zero, Int.alternating_sum_range_choose] theorem sum_powerset_neg_one_pow_card_of_nonempty {α : Type*} {x : Finset α} (h0 : x.Nonempty) : - (∑ m ∈ x.powerset, (-1 : ℤ) ^ m.card) = 0 := by + (∑ m ∈ x.powerset, (-1 : ℤ) ^ #m) = 0 := by classical rw [sum_powerset_neg_one_pow_card] exact if_neg (nonempty_iff_ne_empty.mp h0) diff --git a/Mathlib/Data/Nat/Count.lean b/Mathlib/Data/Nat/Count.lean index cb2e2160ea3aa..3cc54f507cd9c 100644 --- a/Mathlib/Data/Nat/Count.lean +++ b/Mathlib/Data/Nat/Count.lean @@ -36,7 +36,7 @@ theorem count_zero : count p 0 = 0 := by /-- A fintype instance for the set relevant to `Nat.count`. Locally an instance in locale `count` -/ def CountSet.fintype (n : ℕ) : Fintype { i // i < n ∧ p i } := by - apply Fintype.ofFinset ((Finset.range n).filter p) + apply Fintype.ofFinset {x ∈ range n | p x} intro x rw [mem_filter, mem_range] rfl @@ -45,7 +45,7 @@ scoped[Count] attribute [instance] Nat.CountSet.fintype open Count -theorem count_eq_card_filter_range (n : ℕ) : count p n = ((range n).filter p).card := by +theorem count_eq_card_filter_range (n : ℕ) : count p n = #{x ∈ range n | p x} := by rw [count, List.countP_eq_length_filter] rfl @@ -62,7 +62,7 @@ theorem count_monotone : Monotone (count p) := monotone_nat_of_le_succ fun n ↦ by by_cases h : p n <;> simp [count_succ, h] theorem count_add (a b : ℕ) : count p (a + b) = count p a + count (fun k ↦ p (a + k)) b := by - have : Disjoint ((range a).filter p) (((range b).map <| addLeftEmbedding a).filter p) := by + have : Disjoint {x ∈ range a | p x} {x ∈ (range b).map <| addLeftEmbedding a | p x} := by apply disjoint_filter_filter rw [Finset.disjoint_left] simp_rw [mem_map, mem_range, addLeftEmbedding_apply] @@ -114,11 +114,11 @@ theorem count_injective {m n : ℕ} (hm : p m) (hn : p n) (heq : count p m = cou · exact this hn hm heq.symm h.symm (h.lt_or_lt.resolve_left hmn) · simpa [heq] using count_strict_mono hm hmn -theorem count_le_card (hp : (setOf p).Finite) (n : ℕ) : count p n ≤ hp.toFinset.card := by +theorem count_le_card (hp : (setOf p).Finite) (n : ℕ) : count p n ≤ #hp.toFinset := by rw [count_eq_card_filter_range] exact Finset.card_mono fun x hx ↦ hp.mem_toFinset.2 (mem_filter.1 hx).2 -theorem count_lt_card {n : ℕ} (hp : (setOf p).Finite) (hpn : p n) : count p n < hp.toFinset.card := +theorem count_lt_card {n : ℕ} (hp : (setOf p).Finite) (hpn : p n) : count p n < #hp.toFinset := (count_lt_count_succ_iff.2 hpn).trans_le (count_le_card hp _) theorem count_of_forall {n : ℕ} (hp : ∀ n' < n, p n') : count p n = n := by diff --git a/Mathlib/Data/Nat/Factorization/Basic.lean b/Mathlib/Data/Nat/Factorization/Basic.lean index c8c06e616e2d7..f39676359793d 100644 --- a/Mathlib/Data/Nat/Factorization/Basic.lean +++ b/Mathlib/Data/Nat/Factorization/Basic.lean @@ -457,7 +457,7 @@ theorem setOf_pow_dvd_eq_Icc_factorization {n p : ℕ} (pp : p.Prime) (hn : n /-- 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) : - Icc 1 (n.factorization p) = (Ico 1 n).filter fun i : ℕ => p ^ i ∣ n := by + Icc 1 (n.factorization p) = {i ∈ Ico 1 n | p ^ i ∣ n} := by rcases eq_or_ne n 0 with (rfl | hn) · simp ext x @@ -466,11 +466,11 @@ theorem Icc_factorization_eq_pow_dvd (n : ℕ) {p : ℕ} (pp : Prime p) : exact fun _ H => lt_of_le_of_lt H (factorization_lt p hn) theorem factorization_eq_card_pow_dvd (n : ℕ) {p : ℕ} (pp : p.Prime) : - n.factorization p = ((Ico 1 n).filter fun i => p ^ i ∣ n).card := by + n.factorization p = #{i ∈ Ico 1 n | p ^ i ∣ n} := by simp [← Icc_factorization_eq_pow_dvd n pp] theorem Ico_filter_pow_dvd_eq {n p b : ℕ} (pp : p.Prime) (hn : n ≠ 0) (hb : n ≤ p ^ b) : - ((Ico 1 n).filter fun i => p ^ i ∣ n) = (Icc 1 b).filter fun i => p ^ i ∣ n := by + {i ∈ Ico 1 n | p ^ i ∣ n} = {i ∈ Icc 1 b | p ^ i ∣ n} := by ext x simp only [Finset.mem_filter, mem_Ico, mem_Icc, and_congr_left_iff, and_congr_right_iff] rintro h1 - @@ -508,7 +508,7 @@ theorem eq_iff_prime_padicValNat_eq (a b : ℕ) (ha : a ≠ 0) (hb : b ≠ 0) : · simp [factorization_eq_zero_of_non_prime, pp] theorem prod_pow_prime_padicValNat (n : Nat) (hn : n ≠ 0) (m : Nat) (pr : n < m) : - (∏ p ∈ Finset.filter Nat.Prime (Finset.range m), p ^ padicValNat p n) = n := by + ∏ p ∈ range m with p.Prime, p ^ padicValNat p n = n := by -- Porting note: was `nth_rw_rhs` conv => rhs @@ -530,14 +530,14 @@ theorem prod_pow_prime_padicValNat (n : Nat) (hn : n ≠ 0) (m : Nat) (pr : n < -- TODO: Port lemmas from `Data/Nat/Multiplicity` to here, re-written in terms of `factorization` /-- Exactly `n / p` naturals in `[1, n]` are multiples of `p`. See `Nat.card_multiples'` for an alternative spelling of the statement. -/ -theorem card_multiples (n p : ℕ) : card ((Finset.range n).filter fun e => p ∣ e + 1) = n / p := by +theorem card_multiples (n p : ℕ) : #{e ∈ range n | p ∣ e + 1} = n / p := by induction' n with n hn · simp simp [Nat.succ_div, add_ite, add_zero, Finset.range_succ, filter_insert, apply_ite card, card_insert_of_not_mem, hn] /-- Exactly `n / p` naturals in `(0, n]` are multiples of `p`. -/ -theorem Ioc_filter_dvd_card_eq_div (n p : ℕ) : ((Ioc 0 n).filter fun x => p ∣ x).card = n / p := by +theorem Ioc_filter_dvd_card_eq_div (n p : ℕ) : #{x ∈ Ioc 0 n | p ∣ x} = n / p := by induction' n with n IH · simp -- TODO: Golf away `h1` after Yaël PRs a lemma asserting this @@ -550,8 +550,7 @@ theorem Ioc_filter_dvd_card_eq_div (n p : ℕ) : ((Ioc 0 n).filter fun x => p /-- There are exactly `⌊N/n⌋` positive multiples of `n` that are `≤ N`. See `Nat.card_multiples` for a "shifted-by-one" version. -/ -lemma card_multiples' (N n : ℕ) : - ((Finset.range N.succ).filter (fun k ↦ k ≠ 0 ∧ n ∣ k)).card = N / n := by +lemma card_multiples' (N n : ℕ) : #{k ∈ range N.succ | k ≠ 0 ∧ n ∣ k} = N / n := by induction N with | zero => simp [Finset.filter_false_of_mem] | succ N ih => diff --git a/Mathlib/Data/Nat/Factorization/Defs.lean b/Mathlib/Data/Nat/Factorization/Defs.lean index fb64f3dd9d2d6..a274e6adeb506 100644 --- a/Mathlib/Data/Nat/Factorization/Defs.lean +++ b/Mathlib/Data/Nat/Factorization/Defs.lean @@ -163,7 +163,7 @@ theorem factorization_le_iff_dvd {d n : ℕ} (hd : d ≠ 0) (hn : n ≠ 0) : /-- For any `p : ℕ` and any function `g : α → ℕ` that's non-zero on `S : Finset α`, the power of `p` in `S.prod g` equals the sum over `x ∈ S` of the powers of `p` in `g x`. -Generalises `factorization_mul`, which is the special case where `S.card = 2` and `g = id`. -/ +Generalises `factorization_mul`, which is the special case where `#S = 2` and `g = id`. -/ theorem factorization_prod {α : Type*} {S : Finset α} {g : α → ℕ} (hS : ∀ x ∈ S, g x ≠ 0) : (S.prod g).factorization = S.sum fun x => (g x).factorization := by classical diff --git a/Mathlib/Data/Nat/Multiplicity.lean b/Mathlib/Data/Nat/Multiplicity.lean index 61b86c32c88a8..88e818d2af8cd 100644 --- a/Mathlib/Data/Nat/Multiplicity.lean +++ b/Mathlib/Data/Nat/Multiplicity.lean @@ -56,12 +56,12 @@ namespace Nat divides `n`. This set is expressed by filtering `Ico 1 b` where `b` is any bound greater than `log m n`. -/ theorem emultiplicity_eq_card_pow_dvd {m n b : ℕ} (hm : m ≠ 1) (hn : 0 < n) (hb : log m n < b) : - emultiplicity m n = ↑((Finset.Ico 1 b).filter fun i => m ^ i ∣ n).card := + emultiplicity m n = #{i ∈ Ico 1 b | m ^ i ∣ n} := have fin := Nat.multiplicity_finite_iff.2 ⟨hm, hn⟩ calc - emultiplicity m n = ↑(Ico 1 <| multiplicity m n + 1).card := by + emultiplicity m n = #(Ico 1 <| multiplicity m n + 1) := by simp [fin.emultiplicity_eq_multiplicity] - _ = ↑((Finset.Ico 1 b).filter fun i => m ^ i ∣ n).card := + _ = #{i ∈ Ico 1 b | m ^ i ∣ n} := congr_arg _ <| congr_arg card <| Finset.ext fun i => by @@ -107,8 +107,7 @@ theorem emultiplicity_factorial {p : ℕ} (hp : p.Prime) : calc emultiplicity p (n + 1)! = emultiplicity p n ! + emultiplicity p (n + 1) := by rw [factorial_succ, hp.emultiplicity_mul, add_comm] - _ = (∑ i ∈ Ico 1 b, n / p ^ i : ℕ) + - ((Finset.Ico 1 b).filter fun i => p ^ i ∣ n + 1).card := by + _ = (∑ i ∈ Ico 1 b, n / p ^ i : ℕ) + #{i ∈ Ico 1 b | p ^ i ∣ n + 1} := by rw [emultiplicity_factorial hp ((log_mono_right <| le_succ _).trans_lt hb), ← emultiplicity_eq_card_pow_dvd hp.ne_one (succ_pos _) hb] _ = (∑ i ∈ Ico 1 b, (n / p ^ i + if p ^ i ∣ n + 1 then 1 else 0) : ℕ) := by @@ -177,7 +176,7 @@ theorem emultiplicity_factorial_le_div_pred {p : ℕ} (hp : p.Prime) (n : ℕ) : theorem multiplicity_choose_aux {p n b k : ℕ} (hp : p.Prime) (hkn : k ≤ n) : ∑ i ∈ Finset.Ico 1 b, n / p ^ i = ((∑ i ∈ Finset.Ico 1 b, k / p ^ i) + ∑ i ∈ Finset.Ico 1 b, (n - k) / p ^ i) + - ((Finset.Ico 1 b).filter fun i => p ^ i ≤ k % p ^ i + (n - k) % p ^ i).card := + #{i ∈ Ico 1 b | p ^ i ≤ k % p ^ i + (n - k) % p ^ i} := calc ∑ i ∈ Finset.Ico 1 b, n / p ^ i = ∑ i ∈ Finset.Ico 1 b, (k + (n - k)) / p ^ i := by simp only [add_tsub_cancel_of_le hkn] @@ -190,12 +189,10 @@ theorem multiplicity_choose_aux {p n b k : ℕ} (hp : p.Prime) (hkn : k ≤ n) : are added in base `p`. The set is expressed by filtering `Ico 1 b` where `b` is any bound greater than `log p (n + k)`. -/ theorem emultiplicity_choose' {p n k b : ℕ} (hp : p.Prime) (hnb : log p (n + k) < b) : - emultiplicity p (choose (n + k) k) = - ((Ico 1 b).filter fun i => p ^ i ≤ k % p ^ i + n % p ^ i).card := by + emultiplicity p (choose (n + k) k) = #{i ∈ Ico 1 b | p ^ i ≤ k % p ^ i + n % p ^ i} := by have h₁ : emultiplicity p (choose (n + k) k) + emultiplicity p (k ! * n !) = - ((Finset.Ico 1 b).filter fun i => p ^ i ≤ k % p ^ i + n % p ^ i).card + - emultiplicity p (k ! * n !) := by + #{i ∈ Ico 1 b | p ^ i ≤ k % p ^ i + n % p ^ i} + emultiplicity p (k ! * n !) := by rw [← hp.emultiplicity_mul, ← mul_assoc] have := (add_tsub_cancel_right n k) ▸ choose_mul_factorial_mul_factorial (le_add_left k n) rw [this, hp.emultiplicity_factorial hnb, hp.emultiplicity_mul, @@ -211,8 +208,7 @@ theorem emultiplicity_choose' {p n k b : ℕ} (hp : p.Prime) (hnb : log p (n + k are added in base `p`. The set is expressed by filtering `Ico 1 b` where `b` is any bound greater than `log p n`. -/ theorem emultiplicity_choose {p n k b : ℕ} (hp : p.Prime) (hkn : k ≤ n) (hnb : log p n < b) : - emultiplicity p (choose n k) = - ((Ico 1 b).filter fun i => p ^ i ≤ k % p ^ i + (n - k) % p ^ i).card := by + emultiplicity p (choose n k) = #{i ∈ Ico 1 b | p ^ i ≤ k % p ^ i + (n - k) % p ^ i} := by have := Nat.sub_add_cancel hkn convert @emultiplicity_choose' p (n - k) k b hp _ · rw [this] @@ -236,8 +232,8 @@ theorem emultiplicity_choose_prime_pow_add_emultiplicity (hp : p.Prime) (hkn : k le_antisymm (by have hdisj : - Disjoint ((Ico 1 n.succ).filter fun i => p ^ i ≤ k % p ^ i + (p ^ n - k) % p ^ i) - ((Ico 1 n.succ).filter fun i => p ^ i ∣ k) := by + Disjoint {i ∈ Ico 1 n.succ | p ^ i ≤ k % p ^ i + (p ^ n - k) % p ^ i} + {i ∈ Ico 1 n.succ | p ^ i ∣ k} := by simp (config := { contextual := true }) [disjoint_right, *, dvd_iff_mod_eq_zero, Nat.mod_lt _ (pow_pos hp.pos _)] rw [emultiplicity_choose hp hkn (lt_succ_self _), diff --git a/Mathlib/Data/Nat/Nth.lean b/Mathlib/Data/Nat/Nth.lean index 13b3543cb9adc..abe7621d093a2 100644 --- a/Mathlib/Data/Nat/Nth.lean +++ b/Mathlib/Data/Nat/Nth.lean @@ -58,40 +58,40 @@ variable {p} -/ -theorem nth_of_card_le (hf : (setOf p).Finite) {n : ℕ} (hn : hf.toFinset.card ≤ n) : +theorem nth_of_card_le (hf : (setOf p).Finite) {n : ℕ} (hn : #hf.toFinset ≤ n) : nth p n = 0 := by rw [nth, dif_pos hf, List.getD_eq_default]; rwa [Finset.length_sort] theorem nth_eq_getD_sort (h : (setOf p).Finite) (n : ℕ) : nth p n = (h.toFinset.sort (· ≤ ·)).getD n 0 := dif_pos h -theorem nth_eq_orderEmbOfFin (hf : (setOf p).Finite) {n : ℕ} (hn : n < hf.toFinset.card) : +theorem nth_eq_orderEmbOfFin (hf : (setOf p).Finite) {n : ℕ} (hn : n < #hf.toFinset) : nth p n = hf.toFinset.orderEmbOfFin rfl ⟨n, hn⟩ := by rw [nth_eq_getD_sort hf, Finset.orderEmbOfFin_apply, List.getD_eq_getElem, Fin.getElem_fin] theorem nth_strictMonoOn (hf : (setOf p).Finite) : - StrictMonoOn (nth p) (Set.Iio hf.toFinset.card) := by + StrictMonoOn (nth p) (Set.Iio #hf.toFinset) := by rintro m (hm : m < _) n (hn : n < _) h simp only [nth_eq_orderEmbOfFin, *] exact OrderEmbedding.strictMono _ h theorem nth_lt_nth_of_lt_card (hf : (setOf p).Finite) {m n : ℕ} (h : m < n) - (hn : n < hf.toFinset.card) : nth p m < nth p n := + (hn : n < #hf.toFinset) : nth p m < nth p n := nth_strictMonoOn hf (h.trans hn) hn h theorem nth_le_nth_of_lt_card (hf : (setOf p).Finite) {m n : ℕ} (h : m ≤ n) - (hn : n < hf.toFinset.card) : nth p m ≤ nth p n := + (hn : n < #hf.toFinset) : nth p m ≤ nth p n := (nth_strictMonoOn hf).monotoneOn (h.trans_lt hn) hn h theorem lt_of_nth_lt_nth_of_lt_card (hf : (setOf p).Finite) {m n : ℕ} (h : nth p m < nth p n) - (hm : m < hf.toFinset.card) : m < n := + (hm : m < #hf.toFinset) : m < n := not_le.1 fun hle => h.not_le <| nth_le_nth_of_lt_card hf hle hm theorem le_of_nth_le_nth_of_lt_card (hf : (setOf p).Finite) {m n : ℕ} (h : nth p m ≤ nth p n) - (hm : m < hf.toFinset.card) : m ≤ n := + (hm : m < #hf.toFinset) : m ≤ n := not_lt.1 fun hlt => h.not_lt <| nth_lt_nth_of_lt_card hf hlt hm -theorem nth_injOn (hf : (setOf p).Finite) : (Set.Iio hf.toFinset.card).InjOn (nth p) := +theorem nth_injOn (hf : (setOf p).Finite) : (Set.Iio #hf.toFinset).InjOn (nth p) := (nth_strictMonoOn hf).injOn theorem range_nth_of_finite (hf : (setOf p).Finite) : Set.range (nth p) = insert 0 (setOf p) := by @@ -99,20 +99,20 @@ theorem range_nth_of_finite (hf : (setOf p).Finite) : Set.range (nth p) = insert Set.Finite.mem_toFinset] using Set.range_list_getD (hf.toFinset.sort (· ≤ ·)) 0 @[simp] -theorem image_nth_Iio_card (hf : (setOf p).Finite) : nth p '' Set.Iio hf.toFinset.card = setOf p := +theorem image_nth_Iio_card (hf : (setOf p).Finite) : nth p '' Set.Iio #hf.toFinset = setOf p := calc - nth p '' Set.Iio hf.toFinset.card = Set.range (hf.toFinset.orderEmbOfFin rfl) := by + nth p '' Set.Iio #hf.toFinset = Set.range (hf.toFinset.orderEmbOfFin rfl) := by ext x simp only [Set.mem_image, Set.mem_range, Fin.exists_iff, ← nth_eq_orderEmbOfFin hf, Set.mem_Iio, exists_prop] _ = setOf p := by rw [range_orderEmbOfFin, Set.Finite.coe_toFinset] -theorem nth_mem_of_lt_card {n : ℕ} (hf : (setOf p).Finite) (hlt : n < hf.toFinset.card) : +theorem nth_mem_of_lt_card {n : ℕ} (hf : (setOf p).Finite) (hlt : n < #hf.toFinset) : p (nth p n) := (image_nth_Iio_card hf).subset <| Set.mem_image_of_mem _ hlt theorem exists_lt_card_finite_nth_eq (hf : (setOf p).Finite) {x} (h : p x) : - ∃ n, n < hf.toFinset.card ∧ nth p n = x := by + ∃ n, n < #hf.toFinset ∧ nth p n = x := by rwa [← @Set.mem_setOf_eq _ _ p, ← image_nth_Iio_card hf] at h /-! @@ -158,7 +158,7 @@ theorem nth_mem_of_infinite (hf : (setOf p).Infinite) (n : ℕ) : p (nth p n) := -/ theorem exists_lt_card_nth_eq {x} (h : p x) : - ∃ n, (∀ hf : (setOf p).Finite, n < hf.toFinset.card) ∧ nth p n = x := by + ∃ n, (∀ hf : (setOf p).Finite, n < #hf.toFinset) ∧ nth p n = x := by refine (setOf p).finite_or_infinite.elim (fun hf => ?_) fun hf => ?_ · rcases exists_lt_card_finite_nth_eq hf h with ⟨n, hn, hx⟩ exact ⟨n, fun _ => hn, hx⟩ @@ -174,32 +174,32 @@ theorem range_nth_subset : Set.range (nth p) ⊆ insert 0 (setOf p) := (setOf p).finite_or_infinite.elim (fun h => (range_nth_of_finite h).subset) fun h => (range_nth_of_infinite h).trans_subset (Set.subset_insert _ _) -theorem nth_mem (n : ℕ) (h : ∀ hf : (setOf p).Finite, n < hf.toFinset.card) : p (nth p n) := +theorem nth_mem (n : ℕ) (h : ∀ hf : (setOf p).Finite, n < #hf.toFinset) : p (nth p n) := (setOf p).finite_or_infinite.elim (fun hf => nth_mem_of_lt_card hf (h hf)) fun h => nth_mem_of_infinite h n -theorem nth_lt_nth' {m n : ℕ} (hlt : m < n) (h : ∀ hf : (setOf p).Finite, n < hf.toFinset.card) : +theorem nth_lt_nth' {m n : ℕ} (hlt : m < n) (h : ∀ hf : (setOf p).Finite, n < #hf.toFinset) : nth p m < nth p n := (setOf p).finite_or_infinite.elim (fun hf => nth_lt_nth_of_lt_card hf hlt (h _)) fun hf => (nth_lt_nth hf).2 hlt -theorem nth_le_nth' {m n : ℕ} (hle : m ≤ n) (h : ∀ hf : (setOf p).Finite, n < hf.toFinset.card) : +theorem nth_le_nth' {m n : ℕ} (hle : m ≤ n) (h : ∀ hf : (setOf p).Finite, n < #hf.toFinset) : nth p m ≤ nth p n := (setOf p).finite_or_infinite.elim (fun hf => nth_le_nth_of_lt_card hf hle (h _)) fun hf => (nth_le_nth hf).2 hle -theorem le_nth {n : ℕ} (h : ∀ hf : (setOf p).Finite, n < hf.toFinset.card) : n ≤ nth p n := +theorem le_nth {n : ℕ} (h : ∀ hf : (setOf p).Finite, n < #hf.toFinset) : n ≤ nth p n := (setOf p).finite_or_infinite.elim (fun hf => ((nth_strictMonoOn hf).mono <| Set.Iic_subset_Iio.2 (h _)).Iic_id_le _ le_rfl) fun hf => (nth_strictMono hf).id_le _ -theorem isLeast_nth {n} (h : ∀ hf : (setOf p).Finite, n < hf.toFinset.card) : +theorem isLeast_nth {n} (h : ∀ hf : (setOf p).Finite, n < #hf.toFinset) : IsLeast {i | p i ∧ ∀ k < n, nth p k < i} (nth p n) := ⟨⟨nth_mem n h, fun _k hk => nth_lt_nth' hk h⟩, fun _x hx => let ⟨k, hk, hkx⟩ := exists_lt_card_nth_eq hx.1 (lt_or_le k n).elim (fun hlt => absurd hkx (hx.2 _ hlt).ne) fun hle => hkx ▸ nth_le_nth' hle hk⟩ -theorem isLeast_nth_of_lt_card {n : ℕ} (hf : (setOf p).Finite) (hn : n < hf.toFinset.card) : +theorem isLeast_nth_of_lt_card {n : ℕ} (hf : (setOf p).Finite) (hn : n < #hf.toFinset) : IsLeast {i | p i ∧ ∀ k < n, nth p k < i} (nth p n) := isLeast_nth fun _ => hn @@ -210,9 +210,9 @@ theorem isLeast_nth_of_infinite (hf : (setOf p).Infinite) (n : ℕ) : /-- An alternative recursive definition of `Nat.nth`: `Nat.nth s n` is the infimum of `x ∈ s` such that `Nat.nth s k < x` for all `k < n`, if this set is nonempty. We do not assume that the set is nonempty because we use the same "garbage value" `0` both for `sInf` on `ℕ` and for `Nat.nth s n` -for `n ≥ card s`. -/ +for `n ≥ #s`. -/ theorem nth_eq_sInf (p : ℕ → Prop) (n : ℕ) : nth p n = sInf {x | p x ∧ ∀ k < n, nth p k < x} := by - by_cases hn : ∀ hf : (setOf p).Finite, n < hf.toFinset.card + by_cases hn : ∀ hf : (setOf p).Finite, n < #hf.toFinset · exact (isLeast_nth hn).csInf_eq.symm · push_neg at hn rcases hn with ⟨hf, hn⟩ @@ -230,7 +230,7 @@ theorem nth_zero_of_exists [DecidablePred p] (h : ∃ n, p n) : nth p 0 = Nat.fi rw [nth_zero]; convert Nat.sInf_def h theorem nth_eq_zero {n} : - nth p n = 0 ↔ p 0 ∧ n = 0 ∨ ∃ hf : (setOf p).Finite, hf.toFinset.card ≤ n := by + nth p n = 0 ↔ p 0 ∧ n = 0 ∨ ∃ hf : (setOf p).Finite, #hf.toFinset ≤ n := by refine ⟨fun h => ?_, ?_⟩ · simp only [or_iff_not_imp_right, not_exists, not_le] exact fun hn => ⟨h ▸ nth_mem _ hn, nonpos_iff_eq_zero.1 <| h ▸ le_nth hn⟩ @@ -244,7 +244,7 @@ theorem nth_eq_zero_mono (h₀ : ¬p 0) {a b : ℕ} (hab : a ≤ b) (ha : nth p theorem le_nth_of_lt_nth_succ {k a : ℕ} (h : a < nth p (k + 1)) (ha : p a) : a ≤ nth p k := by cases' (setOf p).finite_or_infinite with hf hf · rcases exists_lt_card_finite_nth_eq hf ha with ⟨n, hn, rfl⟩ - cases' lt_or_le (k + 1) hf.toFinset.card with hk hk + cases' lt_or_le (k + 1) #hf.toFinset with hk hk · rwa [(nth_strictMonoOn hf).lt_iff_lt hn hk, Nat.lt_succ_iff, ← (nth_strictMonoOn hf).le_iff_le hn (k.lt_succ_self.trans hk)] at h · rw [nth_of_card_le _ hk] at h @@ -262,7 +262,7 @@ theorem count_nth_zero : count p (nth p 0) = 0 := by exact fun n h₁ h₂ => (mem_range.1 h₁).not_le (Nat.sInf_le h₂) theorem filter_range_nth_subset_insert (k : ℕ) : - (range (nth p (k + 1))).filter p ⊆ insert (nth p k) ((range (nth p k)).filter p) := by + {n ∈ range (nth p (k + 1)) | p n} ⊆ insert (nth p k) {n ∈ range (nth p k) | p n} := by intro a ha simp only [mem_insert, mem_filter, mem_range] at ha ⊢ exact (le_nth_of_lt_nth_succ ha.1 ha.2).eq_or_lt.imp_right fun h => ⟨h, ha.2⟩ @@ -270,8 +270,8 @@ theorem filter_range_nth_subset_insert (k : ℕ) : variable {p} theorem filter_range_nth_eq_insert {k : ℕ} - (hlt : ∀ hf : (setOf p).Finite, k + 1 < hf.toFinset.card) : - (range (nth p (k + 1))).filter p = insert (nth p k) ((range (nth p k)).filter p) := by + (hlt : ∀ hf : (setOf p).Finite, k + 1 < #hf.toFinset) : + {n ∈ range (nth p (k + 1)) | p n} = insert (nth p k) {n ∈ range (nth p k) | p n} := by refine (filter_range_nth_subset_insert p k).antisymm fun a ha => ?_ simp only [mem_insert, mem_filter, mem_range] at ha ⊢ have : nth p k < nth p (k + 1) := nth_lt_nth' k.lt_succ_self hlt @@ -280,15 +280,15 @@ theorem filter_range_nth_eq_insert {k : ℕ} · exact ⟨hlt.trans this, hpa⟩ theorem filter_range_nth_eq_insert_of_finite (hf : (setOf p).Finite) {k : ℕ} - (hlt : k + 1 < hf.toFinset.card) : - (range (nth p (k + 1))).filter p = insert (nth p k) ((range (nth p k)).filter p) := + (hlt : k + 1 < #hf.toFinset) : + {n ∈ range (nth p (k + 1)) | p n} = insert (nth p k) {n ∈ range (nth p k) | p n} := filter_range_nth_eq_insert fun _ => hlt theorem filter_range_nth_eq_insert_of_infinite (hp : (setOf p).Infinite) (k : ℕ) : - (range (nth p (k + 1))).filter p = insert (nth p k) ((range (nth p k)).filter p) := + {n ∈ range (nth p (k + 1)) | p n} = insert (nth p k) {n ∈ range (nth p k) | p n} := filter_range_nth_eq_insert fun hf => absurd hf hp -theorem count_nth {n : ℕ} (hn : ∀ hf : (setOf p).Finite, n < hf.toFinset.card) : +theorem count_nth {n : ℕ} (hn : ∀ hf : (setOf p).Finite, n < #hf.toFinset) : count p (nth p n) = n := by induction' n with k ihk · exact count_nth_zero _ @@ -296,7 +296,7 @@ theorem count_nth {n : ℕ} (hn : ∀ hf : (setOf p).Finite, n < hf.toFinset.car count_eq_card_filter_range, ihk fun hf => lt_of_succ_lt (hn hf)] simp -theorem count_nth_of_lt_card_finite {n : ℕ} (hp : (setOf p).Finite) (hlt : n < hp.toFinset.card) : +theorem count_nth_of_lt_card_finite {n : ℕ} (hp : (setOf p).Finite) (hlt : n < #hp.toFinset) : count p (nth p n) = n := count_nth fun _ => hlt @@ -307,12 +307,12 @@ theorem surjective_count_of_infinite_setOf (h : {n | p n}.Infinite) : Function.Surjective (Nat.count p) := fun n => ⟨nth p n, count_nth_of_infinite h n⟩ -theorem count_nth_succ {n : ℕ} (hn : ∀ hf : (setOf p).Finite, n < hf.toFinset.card) : +theorem count_nth_succ {n : ℕ} (hn : ∀ hf : (setOf p).Finite, n < #hf.toFinset) : count p (nth p n + 1) = n + 1 := by rw [count_succ, count_nth hn, if_pos (nth_mem _ hn)] @[simp] theorem nth_count {n : ℕ} (hpn : p n) : nth p (count p n) = n := - have : ∀ hf : (setOf p).Finite, count p n < hf.toFinset.card := fun hf => count_lt_card hf hpn + have : ∀ hf : (setOf p).Finite, count p n < #hf.toFinset := fun hf => count_lt_card hf hpn count_injective (nth_mem _ this) hpn (count_nth this) theorem nth_lt_of_lt_count {n k : ℕ} (h : k < count p n) : nth p k < n := by diff --git a/Mathlib/Data/Nat/Squarefree.lean b/Mathlib/Data/Nat/Squarefree.lean index 8e1569f711884..40e3df63b05fc 100644 --- a/Mathlib/Data/Nat/Squarefree.lean +++ b/Mathlib/Data/Nat/Squarefree.lean @@ -245,14 +245,14 @@ theorem squarefree_two : Squarefree 2 := by rw [squarefree_iff_nodup_primeFactorsList] <;> simp theorem divisors_filter_squarefree_of_squarefree {n : ℕ} (hn : Squarefree n) : - n.divisors.filter Squarefree = n.divisors := + {d ∈ n.divisors | Squarefree d} = n.divisors := Finset.ext fun d => ⟨@Finset.filter_subset _ _ _ _ d, fun hd => Finset.mem_filter.mpr ⟨hd, hn.squarefree_of_dvd (Nat.dvd_of_mem_divisors hd) ⟩⟩ open UniqueFactorizationMonoid theorem divisors_filter_squarefree {n : ℕ} (h0 : n ≠ 0) : - (n.divisors.filter Squarefree).val = + {d ∈ n.divisors | Squarefree d}.val = (UniqueFactorizationMonoid.normalizedFactors n).toFinset.powerset.val.map fun x => x.val.prod := by rw [(Finset.nodup _).ext ((Finset.nodup _).map_on _)] @@ -303,7 +303,7 @@ theorem divisors_filter_squarefree {n : ℕ} (h0 : n ≠ 0) : theorem sum_divisors_filter_squarefree {n : ℕ} (h0 : n ≠ 0) {α : Type*} [AddCommMonoid α] {f : ℕ → α} : - ∑ i ∈ n.divisors.filter Squarefree, f i = + ∑ d ∈ n.divisors with Squarefree d, f d = ∑ i ∈ (UniqueFactorizationMonoid.normalizedFactors n).toFinset.powerset, f i.val.prod := by rw [Finset.sum_eq_multiset_sum, divisors_filter_squarefree h0, Multiset.map_map, Finset.sum_eq_multiset_sum] @@ -312,7 +312,7 @@ theorem sum_divisors_filter_squarefree {n : ℕ} (h0 : n ≠ 0) {α : Type*} [Ad theorem sq_mul_squarefree_of_pos {n : ℕ} (hn : 0 < n) : ∃ a b : ℕ, 0 < a ∧ 0 < b ∧ b ^ 2 * a = n ∧ Squarefree a := by classical -- Porting note: This line is not needed in Lean 3 - set S := (Finset.range (n + 1)).filter (fun s => s ∣ n ∧ ∃ x, s = x ^ 2) + set S := {s ∈ range (n + 1) | s ∣ n ∧ ∃ x, s = x ^ 2} have hSne : S.Nonempty := by use 1 have h1 : 0 < n ∧ ∃ x : ℕ, 1 = x ^ 2 := ⟨hn, ⟨1, (one_pow 2).symm⟩⟩ diff --git a/Mathlib/Data/Nat/Totient.lean b/Mathlib/Data/Nat/Totient.lean index 59878b2a1d4a9..47f0879f6ad0d 100644 --- a/Mathlib/Data/Nat/Totient.lean +++ b/Mathlib/Data/Nat/Totient.lean @@ -29,8 +29,7 @@ namespace Nat /-- Euler's totient function. This counts the number of naturals strictly less than `n` which are coprime with `n`. -/ -def totient (n : ℕ) : ℕ := - ((range n).filter n.Coprime).card +def totient (n : ℕ) : ℕ := #{a ∈ range n | n.Coprime a} @[inherit_doc] scoped notation "φ" => Nat.totient @@ -42,12 +41,11 @@ theorem totient_zero : φ 0 = 0 := @[simp] theorem totient_one : φ 1 = 1 := rfl -theorem totient_eq_card_coprime (n : ℕ) : φ n = ((range n).filter n.Coprime).card := - rfl +theorem totient_eq_card_coprime (n : ℕ) : φ n = #{a ∈ range n | n.Coprime a} := rfl /-- A characterisation of `Nat.totient` that avoids `Finset`. -/ theorem totient_eq_card_lt_and_coprime (n : ℕ) : φ n = Nat.card { m | m < n ∧ n.Coprime m } := by - let e : { m | m < n ∧ n.Coprime m } ≃ Finset.filter n.Coprime (Finset.range n) := + let e : { m | m < n ∧ n.Coprime m } ≃ {x ∈ range n | n.Coprime x} := { toFun := fun m => ⟨m, by simpa only [Finset.mem_filter, Finset.mem_range] using m.property⟩ invFun := fun m => ⟨m, by simpa only [Finset.mem_filter, Finset.mem_range] using m.property⟩ left_inv := fun m => by simp only [Subtype.coe_mk, Subtype.coe_eta] @@ -70,12 +68,12 @@ theorem totient_eq_zero : ∀ {n : ℕ}, φ n = 0 ↔ n = 0 @[simp] theorem totient_pos {n : ℕ} : 0 < φ n ↔ 0 < n := by simp [pos_iff_ne_zero] theorem filter_coprime_Ico_eq_totient (a n : ℕ) : - ((Ico n (n + a)).filter (Coprime a)).card = totient a := by + #{x ∈ Ico n (n + a) | a.Coprime x} = totient a := by rw [totient, filter_Ico_card_eq_of_periodic, count_eq_card_filter_range] exact periodic_coprime a theorem Ico_filter_coprime_le {a : ℕ} (k n : ℕ) (a_pos : 0 < a) : - ((Ico k (k + n)).filter (Coprime a)).card ≤ totient a * (n / a + 1) := by + #{x ∈ Ico k (k + n) | a.Coprime x} ≤ totient a * (n / a + 1) := by conv_lhs => rw [← Nat.mod_add_div n a] induction' n / a with i ih · rw [← filter_coprime_Ico_eq_totient a k] @@ -86,14 +84,15 @@ theorem Ico_filter_coprime_le {a : ℕ} (k n : ℕ) (a_pos : 0 < a) : simp only [mul_succ] simp_rw [← add_assoc] at ih ⊢ calc - (filter a.Coprime (Ico k (k + n % a + a * i + a))).card = (filter a.Coprime - (Ico k (k + n % a + a * i) ∪ Ico (k + n % a + a * i) (k + n % a + a * i + a))).card := by + #{x ∈ Ico k (k + n % a + a * i + a) | a.Coprime x} + = #{x ∈ Ico k (k + n % a + a * i) ∪ + Ico (k + n % a + a * i) (k + n % a + a * i + a) | a.Coprime x} := by congr rw [Ico_union_Ico_eq_Ico] · rw [add_assoc] exact le_self_add exact le_self_add - _ ≤ (filter a.Coprime (Ico k (k + n % a + a * i))).card + a.totient := by + _ ≤ #{x ∈ Ico k (k + n % a + a * i) | a.Coprime x} + a.totient := by rw [filter_union, ← filter_coprime_Ico_eq_totient a (k + n % a + a * i)] apply card_union_le _ ≤ a.totient * i + a.totient + a.totient := add_le_add_right ih (totient a) @@ -136,7 +135,7 @@ theorem totient_mul {m n : ℕ} (h : m.Coprime n) : φ (m * n) = φ m * φ n := /-- For `d ∣ n`, the totient of `n/d` equals the number of values `k < n` such that `gcd n k = d` -/ theorem totient_div_of_dvd {n d : ℕ} (hnd : d ∣ n) : - φ (n / d) = (filter (fun k : ℕ => n.gcd k = d) (range n)).card := by + φ (n / d) = #{k ∈ range n | n.gcd k = d} := by rcases d.eq_zero_or_pos with (rfl | hd0); · simp [eq_zero_of_zero_dvd hnd] rcases hnd with ⟨x, rfl⟩ rw [Nat.mul_div_cancel_left x hd0] @@ -158,14 +157,14 @@ theorem sum_totient (n : ℕ) : n.divisors.sum φ = n := by rcases n.eq_zero_or_pos with (rfl | hn) · simp rw [← sum_div_divisors n φ] - have : n = ∑ d ∈ n.divisors, (filter (fun k : ℕ => n.gcd k = d) (range n)).card := by + have : n = ∑ d ∈ n.divisors, #{k ∈ range n | n.gcd k = d} := by nth_rw 1 [← card_range n] refine card_eq_sum_card_fiberwise fun x _ => mem_divisors.2 ⟨?_, hn.ne'⟩ apply gcd_dvd_left nth_rw 3 [this] exact sum_congr rfl fun x hx => totient_div_of_dvd (dvd_of_mem_divisors hx) -theorem sum_totient' (n : ℕ) : (∑ m ∈ (range n.succ).filter (· ∣ n), φ m) = n := by +theorem sum_totient' (n : ℕ) : ∑ m ∈ range n.succ with m ∣ n, φ m = n := by convert sum_totient _ using 1 simp only [Nat.divisors, sum_filter, range_eq_Ico] rw [sum_eq_sum_Ico_succ_bot] <;> simp @@ -173,10 +172,10 @@ theorem sum_totient' (n : ℕ) : (∑ m ∈ (range n.succ).filter (· ∣ n), φ /-- When `p` is prime, then the totient of `p ^ (n + 1)` is `p ^ n * (p - 1)` -/ theorem totient_prime_pow_succ {p : ℕ} (hp : p.Prime) (n : ℕ) : φ (p ^ (n + 1)) = p ^ n * (p - 1) := calc - φ (p ^ (n + 1)) = ((range (p ^ (n + 1))).filter (Coprime (p ^ (n + 1)))).card := + φ (p ^ (n + 1)) = #{a ∈ range (p ^ (n + 1)) | (p ^ (n + 1)).Coprime a} := totient_eq_card_coprime _ - _ = (range (p ^ (n + 1)) \ (range (p ^ n)).image (· * p)).card := - (congr_arg card + _ = #(range (p ^ (n + 1)) \ (range (p ^ n)).image (· * p)) := + congr_arg card (by rw [sdiff_eq_filter] apply filter_congr @@ -188,7 +187,7 @@ theorem totient_prime_pow_succ {p : ℕ} (hp : p.Prime) (n : ℕ) : φ (p ^ (n + exact hap (dvd_mul_left _ _) · rintro h ⟨b, rfl⟩ rw [pow_succ'] at ha - exact h b ⟨lt_of_mul_lt_mul_left ha (zero_le _), mul_comm _ _⟩)) + exact h b ⟨lt_of_mul_lt_mul_left ha (zero_le _), mul_comm _ _⟩) _ = _ := by have h1 : Function.Injective (· * p) := mul_left_injective₀ hp.ne_zero have h2 : (range (p ^ n)).image (· * p) ⊆ range (p ^ (n + 1)) := fun a => by diff --git a/Mathlib/Data/PNat/Interval.lean b/Mathlib/Data/PNat/Interval.lean index f10b5ea85500c..c4b3b23d18a7e 100644 --- a/Mathlib/Data/PNat/Interval.lean +++ b/Mathlib/Data/PNat/Interval.lean @@ -52,23 +52,23 @@ theorem map_subtype_embedding_uIcc : (uIcc a b).map (Embedding.subtype _) = uIcc map_subtype_embedding_Icc _ _ @[simp] -theorem card_Icc : (Icc a b).card = b + 1 - a := by +theorem card_Icc : #(Icc a b) = b + 1 - a := by rw [← Nat.card_Icc, ← map_subtype_embedding_Icc, card_map] @[simp] -theorem card_Ico : (Ico a b).card = b - a := by +theorem card_Ico : #(Ico a b) = b - a := by rw [← Nat.card_Ico, ← map_subtype_embedding_Ico, card_map] @[simp] -theorem card_Ioc : (Ioc a b).card = b - a := by +theorem card_Ioc : #(Ioc a b) = b - a := by rw [← Nat.card_Ioc, ← map_subtype_embedding_Ioc, card_map] @[simp] -theorem card_Ioo : (Ioo a b).card = b - a - 1 := by +theorem card_Ioo : #(Ioo a b) = b - a - 1 := by rw [← Nat.card_Ioo, ← map_subtype_embedding_Ioo, card_map] @[simp] -theorem card_uIcc : (uIcc a b).card = (b - a : ℤ).natAbs + 1 := by +theorem card_uIcc : #(uIcc a b) = (b - a : ℤ).natAbs + 1 := by rw [← Nat.card_uIcc, ← map_subtype_embedding_uIcc, card_map] -- Porting note: `simpNF` says `simp` can prove this diff --git a/Mathlib/Data/Pi/Interval.lean b/Mathlib/Data/Pi/Interval.lean index c3f501c51ef40..4db37234636e6 100644 --- a/Mathlib/Data/Pi/Interval.lean +++ b/Mathlib/Data/Pi/Interval.lean @@ -34,16 +34,16 @@ variable (a b : ∀ i, α i) theorem Icc_eq : Icc a b = piFinset fun i => Icc (a i) (b i) := rfl -theorem card_Icc : (Icc a b).card = ∏ i, (Icc (a i) (b i)).card := +theorem card_Icc : #(Icc a b) = ∏ i, #(Icc (a i) (b i)) := card_piFinset _ -theorem card_Ico : (Ico a b).card = (∏ i, (Icc (a i) (b i)).card) - 1 := by +theorem card_Ico : #(Ico a b) = ∏ i, #(Icc (a i) (b i)) - 1 := by rw [card_Ico_eq_card_Icc_sub_one, card_Icc] -theorem card_Ioc : (Ioc a b).card = (∏ i, (Icc (a i) (b i)).card) - 1 := by +theorem card_Ioc : #(Ioc a b) = ∏ i, #(Icc (a i) (b i)) - 1 := by rw [card_Ioc_eq_card_Icc_sub_one, card_Icc] -theorem card_Ioo : (Ioo a b).card = (∏ i, (Icc (a i) (b i)).card) - 2 := by +theorem card_Ioo : #(Ioo a b) = ∏ i, #(Icc (a i) (b i)) - 2 := by rw [card_Ioo_eq_card_Icc_sub_two, card_Icc] end LocallyFiniteOrder @@ -55,11 +55,8 @@ instance instLocallyFiniteOrderBot : LocallyFiniteOrderBot (∀ i, α i) := .ofIic _ (fun b => piFinset fun i => Iic (b i)) fun b x => by simp_rw [mem_piFinset, mem_Iic, le_def] -theorem card_Iic : (Iic b).card = ∏ i, (Iic (b i)).card := - card_piFinset _ - -theorem card_Iio : (Iio b).card = (∏ i, (Iic (b i)).card) - 1 := by - rw [card_Iio_eq_card_Iic_sub_one, card_Iic] +lemma card_Iic : #(Iic b) = ∏ i, #(Iic (b i)) := card_piFinset _ +lemma card_Iio : #(Iio b) = ∏ i, #(Iic (b i)) - 1 := by rw [card_Iio_eq_card_Iic_sub_one, card_Iic] end LocallyFiniteOrderBot @@ -70,11 +67,8 @@ instance instLocallyFiniteOrderTop : LocallyFiniteOrderTop (∀ i, α i) := LocallyFiniteOrderTop.ofIci _ (fun a => piFinset fun i => Ici (a i)) fun a x => by simp_rw [mem_piFinset, mem_Ici, le_def] -theorem card_Ici : (Ici a).card = ∏ i, (Ici (a i)).card := - card_piFinset _ - -theorem card_Ioi : (Ioi a).card = (∏ i, (Ici (a i)).card) - 1 := by - rw [card_Ioi_eq_card_Ici_sub_one, card_Ici] +lemma card_Ici : #(Ici a) = ∏ i, #(Ici (a i)) := card_piFinset _ +lemma card_Ioi : #(Ioi a) = ∏ i, #(Ici (a i)) - 1 := by rw [card_Ioi_eq_card_Ici_sub_one, card_Ici] end LocallyFiniteOrderTop end PartialOrder @@ -84,7 +78,7 @@ variable [∀ i, Lattice (α i)] [∀ i, LocallyFiniteOrder (α i)] (a b : ∀ i theorem uIcc_eq : uIcc a b = piFinset fun i => uIcc (a i) (b i) := rfl -theorem card_uIcc : (uIcc a b).card = ∏ i, (uIcc (a i) (b i)).card := card_Icc _ _ +theorem card_uIcc : #(uIcc a b) = ∏ i, #(uIcc (a i) (b i)) := card_Icc _ _ end Lattice end Pi diff --git a/Mathlib/Data/Sigma/Interval.lean b/Mathlib/Data/Sigma/Interval.lean index ed1676aff61dd..cde30276d4e61 100644 --- a/Mathlib/Data/Sigma/Interval.lean +++ b/Mathlib/Data/Sigma/Interval.lean @@ -55,16 +55,16 @@ section variable (a b : Σ i, α i) -theorem card_Icc : (Icc a b).card = if h : a.1 = b.1 then (Icc (h.rec a.2) b.2).card else 0 := +theorem card_Icc : #(Icc a b) = if h : a.1 = b.1 then #(Icc (h.rec a.2) b.2) else 0 := card_sigmaLift (fun _ => Icc) _ _ -theorem card_Ico : (Ico a b).card = if h : a.1 = b.1 then (Ico (h.rec a.2) b.2).card else 0 := +theorem card_Ico : #(Ico a b) = if h : a.1 = b.1 then #(Ico (h.rec a.2) b.2) else 0 := card_sigmaLift (fun _ => Ico) _ _ -theorem card_Ioc : (Ioc a b).card = if h : a.1 = b.1 then (Ioc (h.rec a.2) b.2).card else 0 := +theorem card_Ioc : #(Ioc a b) = if h : a.1 = b.1 then #(Ioc (h.rec a.2) b.2) else 0 := card_sigmaLift (fun _ => Ioc) _ _ -theorem card_Ioo : (Ioo a b).card = if h : a.1 = b.1 then (Ioo (h.rec a.2) b.2).card else 0 := +theorem card_Ioo : #(Ioo a b) = if h : a.1 = b.1 then #(Ioo (h.rec a.2) b.2) else 0 := card_sigmaLift (fun _ => Ioo) _ _ end diff --git a/Mathlib/Data/Sign.lean b/Mathlib/Data/Sign.lean index e5621d92ed798..0911620f20e6a 100644 --- a/Mathlib/Data/Sign.lean +++ b/Mathlib/Data/Sign.lean @@ -485,7 +485,7 @@ because lean4 infers α to live in a different universe u_2 otherwise -/ private theorem exists_signed_sum_aux {α : Type u_1} [DecidableEq α] (s : Finset α) (f : α → ℤ) : ∃ (β : Type u_1) (t : Finset β) (sgn : β → SignType) (g : β → α), (∀ b, g b ∈ s) ∧ - (t.card = ∑ a ∈ s, (f a).natAbs) ∧ + (#t = ∑ a ∈ s, (f a).natAbs) ∧ ∀ a ∈ s, (∑ b ∈ t, if g b = a then (sgn b : ℤ) else 0) = f a := by refine ⟨(Σ _ : { x // x ∈ s }, ℕ), Finset.univ.sigma fun a => range (f a).natAbs, diff --git a/Mathlib/Data/Sym/Card.lean b/Mathlib/Data/Sym/Card.lean index 4b1dff004b78f..9e65f781aa92d 100644 --- a/Mathlib/Data/Sym/Card.lean +++ b/Mathlib/Data/Sym/Card.lean @@ -121,8 +121,8 @@ namespace Sym2 variable [DecidableEq α] -/-- The `diag` of `s : Finset α` is sent on a finset of `Sym2 α` of card `s.card`. -/ -theorem card_image_diag (s : Finset α) : (s.diag.image Sym2.mk).card = s.card := by +/-- The `diag` of `s : Finset α` is sent on a finset of `Sym2 α` of card `#s`. -/ +theorem card_image_diag (s : Finset α) : #(s.diag.image Sym2.mk) = #s := by rw [card_image_of_injOn, diag_card] rintro ⟨x₀, x₁⟩ hx _ _ h cases Sym2.eq.1 h @@ -130,8 +130,7 @@ theorem card_image_diag (s : Finset α) : (s.diag.image Sym2.mk).card = s.card : · simp only [mem_coe, mem_diag] at hx rw [hx.2] -theorem two_mul_card_image_offDiag (s : Finset α) : - 2 * (s.offDiag.image Sym2.mk).card = s.offDiag.card := by +lemma two_mul_card_image_offDiag (s : Finset α) : 2 * #(s.offDiag.image Sym2.mk) = #s.offDiag := by rw [card_eq_sum_card_image (Sym2.mk : α × α → _), sum_const_nat (Sym2.ind _), mul_comm] rintro x y hxy simp_rw [mem_image, mem_offDiag] at hxy @@ -140,7 +139,7 @@ theorem two_mul_card_image_offDiag (s : Finset α) : obtain ⟨hx, hy, hxy⟩ : x ∈ s ∧ y ∈ s ∧ x ≠ y := by cases h <;> refine ⟨‹_›, ‹_›, ?_⟩ <;> [exact ha; exact ha.symm] have hxy' : y ≠ x := hxy.symm - have : (s.offDiag.filter fun z => Sym2.mk z = s(x, y)) = ({(x, y), (y, x)} : Finset _) := by + have : {z ∈ s.offDiag | Sym2.mk z = s(x, y)} = {(x, y), (y, x)} := by ext ⟨x₁, y₁⟩ rw [mem_filter, mem_insert, mem_singleton, Sym2.eq_iff, Prod.mk.inj_iff, Prod.mk.inj_iff, and_iff_right_iff_imp] @@ -150,11 +149,10 @@ theorem two_mul_card_image_offDiag (s : Finset α) : simp only [not_and, Prod.mk.inj_iff, mem_singleton] exact fun _ => hxy' -/-- The `offDiag` of `s : Finset α` is sent on a finset of `Sym2 α` of card `s.offDiag.card / 2`. +/-- The `offDiag` of `s : Finset α` is sent on a finset of `Sym2 α` of card `#s.offDiag / 2`. This is because every element `s(x, y)` of `Sym2 α` not on the diagonal comes from exactly two pairs: `(x, y)` and `(y, x)`. -/ -theorem card_image_offDiag (s : Finset α) : - (s.offDiag.image Sym2.mk).card = s.card.choose 2 := by +theorem card_image_offDiag (s : Finset α) : #(s.offDiag.image Sym2.mk) = (#s).choose 2 := by rw [Nat.choose_two_right, Nat.mul_sub_left_distrib, mul_one, ← offDiag_card, Nat.div_eq_of_eq_mul_right Nat.zero_lt_two (two_mul_card_image_offDiag s).symm] diff --git a/Mathlib/Data/Sym/Sym2.lean b/Mathlib/Data/Sym/Sym2.lean index c3366dd172bb1..8ae04f76e29e4 100644 --- a/Mathlib/Data/Sym/Sym2.lean +++ b/Mathlib/Data/Sym/Sym2.lean @@ -647,7 +647,7 @@ theorem other_invol {a : α} {z : Sym2 α} (ha : a ∈ z) (hb : Mem.other ha ∈ apply other_eq_other' theorem filter_image_mk_isDiag [DecidableEq α] (s : Finset α) : - ((s ×ˢ s).image Sym2.mk).filter IsDiag = s.diag.image Sym2.mk := by + {a ∈ (s ×ˢ s).image Sym2.mk | a.IsDiag} = s.diag.image Sym2.mk := by ext z induction' z simp only [mem_image, mem_diag, exists_prop, mem_filter, Prod.exists, mem_product] @@ -660,8 +660,7 @@ theorem filter_image_mk_isDiag [DecidableEq α] (s : Finset α) : exact ⟨⟨a, a, ⟨ha, ha⟩, rfl⟩, rfl⟩ theorem filter_image_mk_not_isDiag [DecidableEq α] (s : Finset α) : - (((s ×ˢ s).image Sym2.mk).filter fun a : Sym2 α => ¬a.IsDiag) = - s.offDiag.image Sym2.mk := by + {a ∈ (s ×ˢ s).image Sym2.mk | ¬a.IsDiag} = s.offDiag.image Sym2.mk := by ext z induction z simp only [mem_image, mem_offDiag, mem_filter, Prod.exists, mem_product] From 9b02679beb0d1f7c57b12e9f20779c4e323e6dd5 Mon Sep 17 00:00:00 2001 From: Scott Carnahan <128885296+ScottCarnahan@users.noreply.github.com> Date: Tue, 22 Oct 2024 21:35:14 +0000 Subject: [PATCH 305/505] feat (RingTheory/HahnSeries/Summable) : coeff function for a summable family (#16649) This PR introduces a coefficient function for summable families, and adds some basic formal sum results. - [x] depends on: #16871 [single summable family] - [x] depends on: #16872 [Equiv of families] --- Mathlib/RingTheory/HahnSeries/Summable.lean | 20 ++++++++++++++++++++ 1 file changed, 20 insertions(+) diff --git a/Mathlib/RingTheory/HahnSeries/Summable.lean b/Mathlib/RingTheory/HahnSeries/Summable.lean index a4a57ee6c1624..b4e933391ff50 100644 --- a/Mathlib/RingTheory/HahnSeries/Summable.lean +++ b/Mathlib/RingTheory/HahnSeries/Summable.lean @@ -150,6 +150,17 @@ instance : AddCommMonoid (SummableFamily Γ R α) where ext apply add_assoc +/-- The coefficient function of a summable family, as a finsupp on the parameter type. -/ +@[simps] +def coeff (s : SummableFamily Γ R α) (g : Γ) : α →₀ R where + support := (s.finite_co_support g).toFinset + toFun a := (s a).coeff g + mem_support_toFun a := by simp + +@[simp] +theorem coeff_def (s : SummableFamily Γ R α) (a : α) (g : Γ) : s.coeff g a = (s a).coeff g := + rfl + /-- The infinite sum of a `SummableFamily` of Hahn series. -/ def hsum (s : SummableFamily Γ R α) : HahnSeries Γ R where coeff g := ∑ᶠ i, (s i).coeff g @@ -178,6 +189,15 @@ theorem hsum_add {s t : SummableFamily Γ R α} : (s + t).hsum = s.hsum + t.hsum simp only [hsum_coeff, add_coeff, add_apply] exact finsum_add_distrib (s.finite_co_support _) (t.finite_co_support _) +theorem hsum_coeff_eq_sum_of_subset {s : SummableFamily Γ R α} {g : Γ} {t : Finset α} + (h : { a | (s a).coeff g ≠ 0 } ⊆ t) : s.hsum.coeff g = ∑ i ∈ t, (s i).coeff g := by + simp only [hsum_coeff, finsum_eq_sum _ (s.finite_co_support _)] + exact sum_subset (Set.Finite.toFinset_subset.mpr h) (by simp) + +theorem hsum_coeff_eq_sum {s : SummableFamily Γ R α} {g : Γ} : + s.hsum.coeff g = ∑ i ∈ (s.coeff g).support, (s i).coeff g := by + simp only [hsum_coeff, finsum_eq_sum _ (s.finite_co_support _), coeff_support] + /-- The summable family made of a single Hahn series. -/ @[simps] def single (x : HahnSeries Γ R) : SummableFamily Γ R Unit where From c823b5f7d34fdac912cd365ab30817a4d5152b69 Mon Sep 17 00:00:00 2001 From: Hannah Scholz Date: Tue, 22 Oct 2024 22:22:56 +0000 Subject: [PATCH 306/505] feat: biSup, biInf, biUnion and biInter in NoMaxOrder (#17762) Co-authored-by: Eric Wieser --- Mathlib/Data/Set/Lattice.lean | 24 ++++++++++++++++++++ Mathlib/Order/CompleteLattice.lean | 35 ++++++++++++++++++++++++++++++ 2 files changed, 59 insertions(+) diff --git a/Mathlib/Data/Set/Lattice.lean b/Mathlib/Data/Set/Lattice.lean index 2682e7dba7ab8..f85d0821bddfd 100644 --- a/Mathlib/Data/Set/Lattice.lean +++ b/Mathlib/Data/Set/Lattice.lean @@ -1170,6 +1170,30 @@ theorem union_distrib_iInter_right (s : ι → Set α) (t : Set α) : (⋂ i, s theorem union_distrib_iInter₂_right (s : ∀ i, κ i → Set α) (t : Set α) : (⋂ (i) (j), s i j) ∪ t = ⋂ (i) (j), s i j ∪ t := by simp_rw [union_distrib_iInter_right] +lemma biUnion_lt_eq_iUnion [LT α] [NoMaxOrder α] {s : α → Set β} : + ⋃ (n) (m < n), s m = ⋃ n, s n := biSup_lt_eq_iSup + +lemma biUnion_le_eq_iUnion [Preorder α] {s : α → Set β} : + ⋃ (n) (m ≤ n), s m = ⋃ n, s n := biSup_le_eq_iSup + +lemma biInter_lt_eq_iInter [LT α] [NoMaxOrder α] {s : α → Set β} : + ⋂ (n) (m < n), s m = ⋂ (n), s n := biInf_lt_eq_iInf + +lemma biInter_le_eq_iInter [Preorder α] {s : α → Set β} : + ⋂ (n) (m ≤ n), s m = ⋂ (n), s n := biInf_le_eq_iInf + +lemma biUnion_gt_eq_iUnion [LT α] [NoMinOrder α] {s : α → Set β} : + ⋃ (n) (m > n), s m = ⋃ n, s n := biSup_gt_eq_iSup + +lemma biUnion_ge_eq_iUnion [Preorder α] {s : α → Set β} : + ⋃ (n) (m ≥ n), s m = ⋃ n, s n := biSup_ge_eq_iSup + +lemma biInter_gt_eq_iInf [LT α] [NoMinOrder α] {s : α → Set β} : + ⋂ (n) (m > n), s m = ⋂ n, s n := biInf_gt_eq_iInf + +lemma biInter_ge_eq_iInf [Preorder α] {s : α → Set β} : + ⋂ (n) (m ≥ n), s m = ⋂ n, s n := biInf_ge_eq_iInf + section Function /-! ### Lemmas about `Set.MapsTo` diff --git a/Mathlib/Order/CompleteLattice.lean b/Mathlib/Order/CompleteLattice.lean index b2c06f1f0a3d0..e900ae14df1c3 100644 --- a/Mathlib/Order/CompleteLattice.lean +++ b/Mathlib/Order/CompleteLattice.lean @@ -1073,6 +1073,41 @@ theorem inf_biInf {p : ι → Prop} {f : ∀ i, p i → α} {a : α} (h : ∃ i, (a ⊓ ⨅ (i) (h : p i), f i h) = ⨅ (i) (h : p i), a ⊓ f i h := @sup_biSup αᵒᵈ ι _ p f _ h +lemma biSup_lt_eq_iSup {ι : Type*} [LT ι] [NoMaxOrder ι] {f : ι → α} : + ⨆ (i) (j < i), f j = ⨆ i, f i := by + apply le_antisymm + · exact iSup_le fun _ ↦ iSup₂_le fun _ _ ↦ le_iSup _ _ + · refine iSup_le fun j ↦ ?_ + obtain ⟨i, jlt⟩ := exists_gt j + exact le_iSup_of_le i (le_iSup₂_of_le j jlt le_rfl) + +lemma biSup_le_eq_iSup {ι : Type*} [Preorder ι] {f : ι → α} : + ⨆ (i) (j ≤ i), f j = ⨆ i, f i := by + apply le_antisymm + · exact iSup_le fun _ ↦ iSup₂_le fun _ _ ↦ le_iSup _ _ + · exact iSup_le fun j ↦ le_iSup_of_le j (le_iSup₂_of_le j le_rfl le_rfl) + +lemma biInf_lt_eq_iInf {ι : Type*} [LT ι] [NoMaxOrder ι] {f : ι → α} : + ⨅ (i) (j < i), f j = ⨅ i, f i := + biSup_lt_eq_iSup (α := αᵒᵈ) + +lemma biInf_le_eq_iInf {ι : Type*} [Preorder ι] {f : ι → α} : ⨅ (i) (j ≤ i), f j = ⨅ i, f i := + biSup_le_eq_iSup (α := αᵒᵈ) + +lemma biSup_gt_eq_iSup {ι : Type*} [LT ι] [NoMinOrder ι] {f : ι → α} : + ⨆ (i) (j > i), f j = ⨆ i, f i := + biSup_lt_eq_iSup (ι := ιᵒᵈ) + +lemma biSup_ge_eq_iSup {ι : Type*} [Preorder ι] {f : ι → α} : ⨆ (i) (j ≥ i), f j = ⨆ i, f i := + biSup_le_eq_iSup (ι := ιᵒᵈ) + +lemma biInf_gt_eq_iInf {ι : Type*} [LT ι] [NoMinOrder ι] {f : ι → α} : + ⨅ (i) (j > i), f j = ⨅ i, f i := + biInf_lt_eq_iInf (ι := ιᵒᵈ) + +lemma biInf_ge_eq_iInf {ι : Type*} [Preorder ι] {f : ι → α} : ⨅ (i) (j ≥ i), f j = ⨅ i, f i := + biInf_le_eq_iInf (ι := ιᵒᵈ) + /-! ### `iSup` and `iInf` under `Prop` -/ From 421f2fccbb0244342f570fbf62003f22744ef2b1 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Tue, 22 Oct 2024 23:10:55 +0000 Subject: [PATCH 307/505] =?UTF-8?q?chore:=20remove=20C=E2=8B=86-algebra=20?= =?UTF-8?q?import=20from=20`LinearAlgebra.Matrix.PosDef`=20(#18086)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- .../LinearAlgebra/Matrix/HermitianFunctionalCalculus.lean | 2 ++ Mathlib/LinearAlgebra/Matrix/Spectrum.lean | 8 +------- 2 files changed, 3 insertions(+), 7 deletions(-) diff --git a/Mathlib/LinearAlgebra/Matrix/HermitianFunctionalCalculus.lean b/Mathlib/LinearAlgebra/Matrix/HermitianFunctionalCalculus.lean index a1b13157f47c4..70a529e030dc1 100644 --- a/Mathlib/LinearAlgebra/Matrix/HermitianFunctionalCalculus.lean +++ b/Mathlib/LinearAlgebra/Matrix/HermitianFunctionalCalculus.lean @@ -8,6 +8,8 @@ import Mathlib.LinearAlgebra.Matrix.Spectrum import Mathlib.LinearAlgebra.Eigenspace.Matrix import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Unique import Mathlib.Topology.ContinuousMap.Units +import Mathlib.Analysis.Matrix +import Mathlib.Topology.UniformSpace.Matrix /-! # Continuous Functional Calculus for Hermitian Matrices diff --git a/Mathlib/LinearAlgebra/Matrix/Spectrum.lean b/Mathlib/LinearAlgebra/Matrix/Spectrum.lean index e20e0a672cfc3..9f15fd930b639 100644 --- a/Mathlib/LinearAlgebra/Matrix/Spectrum.lean +++ b/Mathlib/LinearAlgebra/Matrix/Spectrum.lean @@ -7,7 +7,6 @@ import Mathlib.Analysis.InnerProductSpace.Spectrum import Mathlib.Data.Matrix.Rank import Mathlib.LinearAlgebra.Matrix.Diagonal import Mathlib.LinearAlgebra.Matrix.Hermitian -import Mathlib.Analysis.CStarAlgebra.Matrix import Mathlib.Topology.Algebra.Module.FiniteDimension /-! # Spectral theory of hermitian matrices @@ -54,12 +53,7 @@ lemma mulVec_eigenvectorBasis (j : n) : /-- The spectrum of a Hermitian matrix `A` coincides with the spectrum of `toEuclideanLin A`. -/ theorem spectrum_toEuclideanLin : spectrum 𝕜 (toEuclideanLin A) = spectrum 𝕜 A := - AlgEquiv.spectrum_eq - (AlgEquiv.trans - ((toEuclideanCLM : Matrix n n 𝕜 ≃⋆ₐ[𝕜] EuclideanSpace 𝕜 n →L[𝕜] EuclideanSpace 𝕜 n) : - Matrix n n 𝕜 ≃ₐ[𝕜] EuclideanSpace 𝕜 n →L[𝕜] EuclideanSpace 𝕜 n) - (Module.End.toContinuousLinearMap (EuclideanSpace 𝕜 n)).symm) - _ + AlgEquiv.spectrum_eq (Matrix.toLinAlgEquiv (PiLp.basisFun 2 𝕜 n)) _ /-- Eigenvalues of a hermitian matrix A are in the ℝ spectrum of A. -/ theorem eigenvalues_mem_spectrum_real (i : n) : hA.eigenvalues i ∈ spectrum ℝ A := by From 69d256b3c87085fec8063376dfb231ee1fd345d8 Mon Sep 17 00:00:00 2001 From: mathlib4-update-dependencies-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Tue, 22 Oct 2024 23:48:07 +0000 Subject: [PATCH 308/505] chore: update Mathlib dependencies 2024-10-22 (#18093) This PR updates the Mathlib dependencies. --- lake-manifest.json | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lake-manifest.json b/lake-manifest.json index a34431f99113a..02b5a77289c00 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -25,7 +25,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "b20a88676fd00affb90cbc9f1ff004ae588103b3", + "rev": "9ac12945862fa39eab7795c2f79bb9aa0c8e332c", "name": "aesop", "manifestFile": "lake-manifest.json", "inputRev": "master", From cd8e93f720b73f9a7ce4d6162300cdfe8bd6896e Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 23 Oct 2024 01:35:03 +0000 Subject: [PATCH 309/505] chore: generalize from `AddCommGroup` to `AddCommMonoid` (#18095) This lets many theorems apply to `M := R` where `R` is a ring. --- Mathlib/LinearAlgebra/Basis/Cardinality.lean | 2 +- Mathlib/LinearAlgebra/FreeModule/Finite/Basic.lean | 2 +- Mathlib/RingTheory/Finiteness.lean | 2 +- 3 files changed, 3 insertions(+), 3 deletions(-) diff --git a/Mathlib/LinearAlgebra/Basis/Cardinality.lean b/Mathlib/LinearAlgebra/Basis/Cardinality.lean index 284ad5ed30aa4..64108a9f48908 100644 --- a/Mathlib/LinearAlgebra/Basis/Cardinality.lean +++ b/Mathlib/LinearAlgebra/Basis/Cardinality.lean @@ -21,7 +21,7 @@ variable {R : Type u} {M : Type v} section Semiring -variable [Semiring R] [AddCommGroup M] [Nontrivial R] [Module R M] +variable [Semiring R] [AddCommMonoid M] [Nontrivial R] [Module R M] -- One might hope that a finite spanning set implies that any linearly independent set is finite. -- While this is true over a division ring diff --git a/Mathlib/LinearAlgebra/FreeModule/Finite/Basic.lean b/Mathlib/LinearAlgebra/FreeModule/Finite/Basic.lean index fd4961dd7d2fc..7cc74f0c2a909 100644 --- a/Mathlib/LinearAlgebra/FreeModule/Finite/Basic.lean +++ b/Mathlib/LinearAlgebra/FreeModule/Finite/Basic.lean @@ -22,7 +22,7 @@ universe u v w /-- If a free module is finite, then the arbitrary basis is finite. -/ noncomputable instance Module.Free.ChooseBasisIndex.fintype (R : Type u) (M : Type v) - [Semiring R] [AddCommGroup M] [Module R M] [Module.Free R M] [Module.Finite R M] : + [Semiring R] [AddCommMonoid M] [Module R M] [Module.Free R M] [Module.Finite R M] : Fintype (Module.Free.ChooseBasisIndex R M) := by refine @Fintype.ofFinite _ ?_ cases subsingleton_or_nontrivial R diff --git a/Mathlib/RingTheory/Finiteness.lean b/Mathlib/RingTheory/Finiteness.lean index 264b1bb8d9350..caa9661139eaf 100644 --- a/Mathlib/RingTheory/Finiteness.lean +++ b/Mathlib/RingTheory/Finiteness.lean @@ -724,7 +724,7 @@ instance Module.Finite.tensorProduct [CommSemiring R] [AddCommMonoid M] [Module out := (TensorProduct.map₂_mk_top_top_eq_top R M N).subst (hM.out.map₂ _ hN.out) /-- If a free module is finite, then any arbitrary basis is finite. -/ -lemma Module.Finite.finite_basis {R M} [Semiring R] [Nontrivial R] [AddCommGroup M] [Module R M] +lemma Module.Finite.finite_basis {R M} [Semiring R] [Nontrivial R] [AddCommMonoid M] [Module R M] {ι} [Module.Finite R M] (b : Basis ι R M) : _root_.Finite ι := let ⟨s, hs⟩ := ‹Module.Finite R M› From ebbced44ffe5f9f6ae1efad8056ba51b2e85eb30 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Wed, 23 Oct 2024 02:08:04 +0000 Subject: [PATCH 310/505] chore(Algebra): re-add `@[simp]` attributes lost in the port (#18057) This PR re-adds the `@[simp]` attributes that had a porting note that they could be proved by `@[simp]`, but where that does not appear to be the case currently. Of course, the `simp` set may have changed sufficiently to require un`@[simp]`ing some of these anyway, but we should add a more up to date comment in that case. --- Mathlib/Algebra/AddTorsor.lean | 2 -- Mathlib/Algebra/Algebra/Equiv.lean | 4 ++-- Mathlib/Algebra/Algebra/NonUnitalHom.lean | 2 +- Mathlib/Algebra/Group/Equiv/Basic.lean | 6 ++++-- Mathlib/Algebra/Group/Pointwise/Set/Basic.lean | 1 - Mathlib/Algebra/Group/Subgroup/Basic.lean | 1 - Mathlib/Algebra/Group/Submonoid/Membership.lean | 8 ++++---- Mathlib/Algebra/Lie/Submodule.lean | 2 +- Mathlib/Algebra/MonoidAlgebra/Defs.lean | 8 ++++++-- Mathlib/Algebra/MvPolynomial/Basic.lean | 12 ++++++------ Mathlib/Algebra/MvPolynomial/CommRing.lean | 4 ++-- Mathlib/Algebra/Order/AbsoluteValue.lean | 8 ++++---- Mathlib/Algebra/PUnitInstances/GCDMonoid.lean | 4 ++-- Mathlib/Algebra/Polynomial/Basic.lean | 4 ++-- Mathlib/Algebra/Polynomial/Derivative.lean | 8 ++++---- Mathlib/Algebra/Polynomial/Laurent.lean | 6 +++--- Mathlib/Algebra/Ring/Equiv.lean | 2 +- Mathlib/Algebra/Ring/Hom/Defs.lean | 3 +-- Mathlib/Algebra/Ring/Int.lean | 6 ++---- Mathlib/Algebra/Ring/Subring/Basic.lean | 4 ++-- Mathlib/Algebra/Star/Subalgebra.lean | 1 - Mathlib/FieldTheory/IntermediateField/Basic.lean | 2 +- 22 files changed, 48 insertions(+), 50 deletions(-) diff --git a/Mathlib/Algebra/AddTorsor.lean b/Mathlib/Algebra/AddTorsor.lean index 38aff47034d89..d615b53f28671 100644 --- a/Mathlib/Algebra/AddTorsor.lean +++ b/Mathlib/Algebra/AddTorsor.lean @@ -165,8 +165,6 @@ namespace Set open Pointwise --- porting note (#10618): simp can prove this ---@[simp] theorem singleton_vsub_self (p : P) : ({p} : Set P) -ᵥ {p} = {(0 : G)} := by rw [Set.singleton_vsub_singleton, vsub_self] diff --git a/Mathlib/Algebra/Algebra/Equiv.lean b/Mathlib/Algebra/Algebra/Equiv.lean index fa867ba948dfb..89b6ce8e80381 100644 --- a/Mathlib/Algebra/Algebra/Equiv.lean +++ b/Mathlib/Algebra/Algebra/Equiv.lean @@ -288,13 +288,13 @@ def symm (e : A₁ ≃ₐ[R] A₂) : A₂ ≃ₐ[R] A₁ := theorem invFun_eq_symm {e : A₁ ≃ₐ[R] A₂} : e.invFun = e.symm := rfl ---@[simp] -- Porting note (#10618): simp can prove this once symm_mk is introduced +@[simp] theorem coe_apply_coe_coe_symm_apply {F : Type*} [EquivLike F A₁ A₂] [AlgEquivClass F R A₁ A₂] (f : F) (x : A₂) : f ((f : A₁ ≃ₐ[R] A₂).symm x) = x := EquivLike.right_inv f x ---@[simp] -- Porting note (#10618): simp can prove this once symm_mk is introduced +@[simp] theorem coe_coe_symm_apply_coe_apply {F : Type*} [EquivLike F A₁ A₂] [AlgEquivClass F R A₁ A₂] (f : F) (x : A₁) : (f : A₁ ≃ₐ[R] A₂).symm (f x) = x := diff --git a/Mathlib/Algebra/Algebra/NonUnitalHom.lean b/Mathlib/Algebra/Algebra/NonUnitalHom.lean index 7eb27b49ee74a..f91df6f1aa69a 100644 --- a/Mathlib/Algebra/Algebra/NonUnitalHom.lean +++ b/Mathlib/Algebra/Algebra/NonUnitalHom.lean @@ -245,7 +245,7 @@ theorem coe_mulHom_mk (f : A →ₛₙₐ[φ] B) (h₁ h₂ h₃ h₄) : ((⟨⟨⟨f, h₁⟩, h₂, h₃⟩, h₄⟩ : A →ₛₙₐ[φ] B) : A →ₙ* B) = ⟨f, h₄⟩ := by rfl --- @[simp] -- Porting note (#10618) : simp can prove this +@[simp] -- Marked as `@[simp]` because `MulActionSemiHomClass.map_smulₛₗ` can't be. protected theorem map_smul (f : A →ₛₙₐ[φ] B) (c : R) (x : A) : f (c • x) = (φ c) • f x := map_smulₛₗ _ _ _ diff --git a/Mathlib/Algebra/Group/Equiv/Basic.lean b/Mathlib/Algebra/Group/Equiv/Basic.lean index cb72fdfe7db40..8d70c2f054093 100644 --- a/Mathlib/Algebra/Group/Equiv/Basic.lean +++ b/Mathlib/Algebra/Group/Equiv/Basic.lean @@ -472,11 +472,13 @@ end Mul section MulOneClass variable [MulOneClass M] [MulOneClass N] [MulOneClass P] --- Porting note (#10618): `simp` can prove this +-- Porting note (#10618): `simp` can prove this but it is a valid `dsimp` lemma. +-- However, we would need to redesign the the `dsimp` set to make this `@[simp]`. @[to_additive] theorem coe_monoidHom_refl : (refl M : M →* M) = MonoidHom.id M := rfl --- Porting note (#10618): `simp` can prove this +-- Porting note (#10618): `simp` can prove this but it is a valid `dsimp` lemma. +-- However, we would need to redesign the the `dsimp` set to make this `@[simp]`. @[to_additive] lemma coe_monoidHom_trans (e₁ : M ≃* N) (e₂ : N ≃* P) : (e₁.trans e₂ : M →* P) = (e₂ : N →* P).comp ↑e₁ := rfl diff --git a/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean b/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean index 946330326b6c8..7dc3e11ce73a6 100644 --- a/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean +++ b/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean @@ -533,7 +533,6 @@ theorem div_singleton : s / {b} = (· / b) '' s := theorem singleton_div : {a} / t = (· / ·) a '' t := image2_singleton_left --- Porting note (#10618): simp can prove this @[to_additive] theorem singleton_div_singleton : ({a} : Set α) / {b} = {a / b} := image2_singleton diff --git a/Mathlib/Algebra/Group/Subgroup/Basic.lean b/Mathlib/Algebra/Group/Subgroup/Basic.lean index 4f1bb54083fb7..e53212f0a17ad 100644 --- a/Mathlib/Algebra/Group/Subgroup/Basic.lean +++ b/Mathlib/Algebra/Group/Subgroup/Basic.lean @@ -1969,7 +1969,6 @@ theorem normalCore_eq_iSup (H : Subgroup G) : theorem normalCore_eq_self (H : Subgroup G) [H.Normal] : H.normalCore = H := le_antisymm H.normalCore_le (normal_le_normalCore.mpr le_rfl) --- @[simp] -- Porting note (#10618): simp can prove this theorem normalCore_idempotent (H : Subgroup G) : H.normalCore.normalCore = H.normalCore := H.normalCore.normalCore_eq_self diff --git a/Mathlib/Algebra/Group/Submonoid/Membership.lean b/Mathlib/Algebra/Group/Submonoid/Membership.lean index 2c5e76300f647..9b396e3f6c38b 100644 --- a/Mathlib/Algebra/Group/Submonoid/Membership.lean +++ b/Mathlib/Algebra/Group/Submonoid/Membership.lean @@ -56,7 +56,7 @@ theorem coe_multiset_prod {M} [CommMonoid M] [SetLike B M] [SubmonoidClass B M] (m.prod : M) = (m.map (↑)).prod := (SubmonoidClass.subtype S : _ →* M).map_multiset_prod m -@[to_additive (attr := norm_cast)] -- Porting note (#10618): removed `simp`, `simp` can prove it +@[to_additive (attr := norm_cast, simp)] theorem coe_finset_prod {ι M} [CommMonoid M] [SetLike B M] [SubmonoidClass B M] (f : ι → S) (s : Finset ι) : ↑(∏ i ∈ s, f i) = (∏ i ∈ s, f i : M) := map_prod (SubmonoidClass.subtype S) f s @@ -97,16 +97,16 @@ namespace Submonoid variable (s : Submonoid M) -@[to_additive (attr := norm_cast)] -- Porting note (#10618): removed `simp`, `simp` can prove it +@[to_additive (attr := norm_cast)] theorem coe_list_prod (l : List s) : (l.prod : M) = (l.map (↑)).prod := map_list_prod s.subtype l -@[to_additive (attr := norm_cast)] -- Porting note (#10618): removed `simp`, `simp` can prove it +@[to_additive (attr := norm_cast)] theorem coe_multiset_prod {M} [CommMonoid M] (S : Submonoid M) (m : Multiset S) : (m.prod : M) = (m.map (↑)).prod := S.subtype.map_multiset_prod m -@[to_additive (attr := norm_cast, simp)] +@[to_additive (attr := norm_cast)] theorem coe_finset_prod {ι M} [CommMonoid M] (S : Submonoid M) (f : ι → S) (s : Finset ι) : ↑(∏ i ∈ s, f i) = (∏ i ∈ s, f i : M) := map_prod S.subtype f s diff --git a/Mathlib/Algebra/Lie/Submodule.lean b/Mathlib/Algebra/Lie/Submodule.lean index 6c64bbfcaee74..f09b3652b769d 100644 --- a/Mathlib/Algebra/Lie/Submodule.lean +++ b/Mathlib/Algebra/Lie/Submodule.lean @@ -83,7 +83,7 @@ instance (priority := mid) coeSubmodule : CoeOut (LieSubmodule R L M) (Submodule theorem coe_toSubmodule : ((N : Submodule R M) : Set M) = N := rfl --- Porting note (#10618): `simp` can prove this after `mem_coeSubmodule` is added to the simp set, +-- `simp` can prove this after `mem_coeSubmodule` is added to the simp set, -- but `dsimp` can't. @[simp, nolint simpNF] theorem mem_carrier {x : M} : x ∈ N.carrier ↔ x ∈ (N : Set M) := diff --git a/Mathlib/Algebra/MonoidAlgebra/Defs.lean b/Mathlib/Algebra/MonoidAlgebra/Defs.lean index 2d212d128bc5d..2b71376c102f0 100644 --- a/Mathlib/Algebra/MonoidAlgebra/Defs.lean +++ b/Mathlib/Algebra/MonoidAlgebra/Defs.lean @@ -755,7 +755,9 @@ protected noncomputable def opRingEquiv [Monoid G] : erw [unop_op, unop_op, single_mul_single] simp } --- @[simp] -- Porting note (#10618): simp can prove this +-- @[simp] -- Porting note (#10618): simp can prove this. +-- More specifically, the LHS simplifies to `Finsupp.single`, which implies there's some +-- defeq abuse going on. theorem opRingEquiv_single [Monoid G] (r : k) (x : G) : MonoidAlgebra.opRingEquiv (op (single x r)) = single (op x) (op r) := by simp @@ -1413,7 +1415,9 @@ protected noncomputable def opRingEquiv [AddCommMonoid G] : erw [mapRange_single, single_mul_single, mapRange_single, mapRange_single] simp only [mapRange_single, single_mul_single, ← op_mul, add_comm] } --- @[simp] -- Porting note (#10618): simp can prove this +-- @[simp] -- Porting note (#10618): simp can prove this. +-- More specifically, the LHS simplifies to `Finsupp.single`, which implies there's some +-- defeq abuse going on. theorem opRingEquiv_single [AddCommMonoid G] (r : k) (x : G) : AddMonoidAlgebra.opRingEquiv (op (single x r)) = single x (op r) := by simp diff --git a/Mathlib/Algebra/MvPolynomial/Basic.lean b/Mathlib/Algebra/MvPolynomial/Basic.lean index 910e409090c04..1e8efa3f5e687 100644 --- a/Mathlib/Algebra/MvPolynomial/Basic.lean +++ b/Mathlib/Algebra/MvPolynomial/Basic.lean @@ -191,10 +191,10 @@ theorem monomial_left_inj {s t : σ →₀ ℕ} {r : R} (hr : r ≠ 0) : theorem C_apply : (C a : MvPolynomial σ R) = monomial 0 a := rfl --- Porting note (#10618): `simp` can prove this +@[simp] theorem C_0 : C 0 = (0 : MvPolynomial σ R) := map_zero _ --- Porting note (#10618): `simp` can prove this +@[simp] theorem C_1 : C 1 = (1 : MvPolynomial σ R) := rfl @@ -203,15 +203,15 @@ theorem C_mul_monomial : C a * monomial s a' = monomial s (a * a') := by show AddMonoidAlgebra.single _ _ * AddMonoidAlgebra.single _ _ = AddMonoidAlgebra.single _ _ simp [C_apply, single_mul_single] --- Porting note (#10618): `simp` can prove this +@[simp] theorem C_add : (C (a + a') : MvPolynomial σ R) = C a + C a' := Finsupp.single_add _ _ _ --- Porting note (#10618): `simp` can prove this +@[simp] theorem C_mul : (C (a * a') : MvPolynomial σ R) = C a * C a' := C_mul_monomial.symm --- Porting note (#10618): `simp` can prove this +@[simp] theorem C_pow (a : R) (n : ℕ) : (C (a ^ n) : MvPolynomial σ R) = C a ^ n := map_pow _ _ _ @@ -303,7 +303,7 @@ theorem C_mul_X_pow_eq_monomial {s : σ} {a : R} {n : ℕ} : theorem C_mul_X_eq_monomial {s : σ} {a : R} : C a * X s = monomial (Finsupp.single s 1) a := by rw [← C_mul_X_pow_eq_monomial, pow_one] --- Porting note (#10618): `simp` can prove this +@[simp] theorem monomial_zero {s : σ →₀ ℕ} : monomial s (0 : R) = 0 := Finsupp.single_zero _ diff --git a/Mathlib/Algebra/MvPolynomial/CommRing.lean b/Mathlib/Algebra/MvPolynomial/CommRing.lean index 7e4cbd57dbbc2..fc473489090b6 100644 --- a/Mathlib/Algebra/MvPolynomial/CommRing.lean +++ b/Mathlib/Algebra/MvPolynomial/CommRing.lean @@ -55,11 +55,11 @@ instance instCommRingMvPolynomial : CommRing (MvPolynomial σ R) := variable (σ a a') --- @[simp] -- Porting note (#10618): simp can prove this +@[simp] theorem C_sub : (C (a - a') : MvPolynomial σ R) = C a - C a' := RingHom.map_sub _ _ _ --- @[simp] -- Porting note (#10618): simp can prove this +@[simp] theorem C_neg : (C (-a) : MvPolynomial σ R) = -C a := RingHom.map_neg _ _ diff --git a/Mathlib/Algebra/Order/AbsoluteValue.lean b/Mathlib/Algebra/Order/AbsoluteValue.lean index fae16b0ffcd6d..8f31d949787d9 100644 --- a/Mathlib/Algebra/Order/AbsoluteValue.lean +++ b/Mathlib/Algebra/Order/AbsoluteValue.lean @@ -93,7 +93,7 @@ protected theorem eq_zero {x : R} : abv x = 0 ↔ x = 0 := protected theorem add_le (x y : R) : abv (x + y) ≤ abv x + abv y := abv.add_le' x y --- Porting note (#10618): was `@[simp]` but `simp` can prove it +@[simp] protected theorem map_mul (x y : R) : abv (x * y) = abv x * abv y := abv.map_mul' x y @@ -113,7 +113,7 @@ protected theorem ne_zero {x : R} (hx : x ≠ 0) : abv x ≠ 0 := theorem map_one_of_isLeftRegular (h : IsLeftRegular (abv 1)) : abv 1 = 1 := h <| by simp [← abv.map_mul] --- Porting note (#10618): was `@[simp]` but `simp` can prove it +@[simp] protected theorem map_zero : abv 0 = 0 := abv.eq_zero.2 rfl @@ -145,7 +145,7 @@ section IsDomain variable {R S : Type*} [Semiring R] [OrderedRing S] (abv : AbsoluteValue R S) variable [IsDomain S] [Nontrivial R] --- Porting note (#10618): was `@[simp]` but `simp` can prove it +@[simp] protected theorem map_one : abv 1 = 1 := abv.map_one_of_isLeftRegular (isRegular_of_ne_zero <| abv.ne_zero one_ne_zero).left @@ -170,7 +170,7 @@ def toMonoidHom : R →* S := theorem coe_toMonoidHom : ⇑abv.toMonoidHom = abv := rfl --- Porting note (#10618): was `@[simp]` but `simp` can prove it +@[simp] protected theorem map_pow (a : R) (n : ℕ) : abv (a ^ n) = abv a ^ n := abv.toMonoidHom.map_pow a n diff --git a/Mathlib/Algebra/PUnitInstances/GCDMonoid.lean b/Mathlib/Algebra/PUnitInstances/GCDMonoid.lean index d80a17643fd74..3609ecaa55b47 100644 --- a/Mathlib/Algebra/PUnitInstances/GCDMonoid.lean +++ b/Mathlib/Algebra/PUnitInstances/GCDMonoid.lean @@ -32,11 +32,11 @@ instance normalizedGCDMonoid : NormalizedGCDMonoid PUnit where normalize_gcd := by intros; rfl normalize_lcm := by intros; rfl --- Porting note (#10618): simpNF lint: simp can prove this @[simp] +@[simp] theorem gcd_eq {x y : PUnit} : gcd x y = unit := rfl --- Porting note (#10618): simpNF lint: simp can prove this @[simp] +@[simp] theorem lcm_eq {x y : PUnit} : lcm x y = unit := rfl diff --git a/Mathlib/Algebra/Polynomial/Basic.lean b/Mathlib/Algebra/Polynomial/Basic.lean index b3f8ed139f76f..f095ae814ba7d 100644 --- a/Mathlib/Algebra/Polynomial/Basic.lean +++ b/Mathlib/Algebra/Polynomial/Basic.lean @@ -379,7 +379,7 @@ theorem toFinsupp_monomial (n : ℕ) (r : R) : (monomial n r).toFinsupp = Finsup theorem ofFinsupp_single (n : ℕ) (r : R) : (⟨Finsupp.single n r⟩ : R[X]) = monomial n r := by simp [monomial] --- @[simp] -- Porting note (#10618): simp can prove this +@[simp] theorem monomial_zero_right (n : ℕ) : monomial n (0 : R) = 0 := (monomial n).map_zero @@ -1028,7 +1028,7 @@ theorem coeff_sub (p q : R[X]) (n : ℕ) : coeff (p - q) n = coeff p n - coeff q -- Porting note: The last rule should be `apply`ed. rw [← ofFinsupp_sub, coeff, coeff, coeff]; apply Finsupp.sub_apply --- @[simp] -- Porting note (#10618): simp can prove this +@[simp] theorem monomial_neg (n : ℕ) (a : R) : monomial n (-a) = -monomial n a := by rw [eq_neg_iff_add_eq_zero, ← monomial_add, neg_add_cancel, monomial_zero_right] diff --git a/Mathlib/Algebra/Polynomial/Derivative.lean b/Mathlib/Algebra/Polynomial/Derivative.lean index 9108354da435c..5e8f7007d1c5b 100644 --- a/Mathlib/Algebra/Polynomial/Derivative.lean +++ b/Mathlib/Algebra/Polynomial/Derivative.lean @@ -71,7 +71,7 @@ theorem coeff_derivative (p : R[X]) (n : ℕ) : push_neg at h simp [h] --- Porting note (#10618): removed `simp`: `simp` can prove it. +@[simp] theorem derivative_zero : derivative (0 : R[X]) = 0 := derivative.map_zero @@ -116,7 +116,7 @@ theorem derivative_X : derivative (X : R[X]) = 1 := theorem derivative_one : derivative (1 : R[X]) = 0 := derivative_C --- Porting note (#10618): removed `simp`: `simp` can prove it. +@[simp] theorem derivative_add {f g : R[X]} : derivative (f + g) = derivative f + derivative g := derivative.map_add f g @@ -560,14 +560,14 @@ section Ring variable [Ring R] --- Porting note (#10618): removed `simp`: `simp` can prove it. +@[simp] theorem derivative_neg (f : R[X]) : derivative (-f) = -derivative f := LinearMap.map_neg derivative f theorem iterate_derivative_neg {f : R[X]} {k : ℕ} : derivative^[k] (-f) = -derivative^[k] f := iterate_map_neg derivative k f --- Porting note (#10618): removed `simp`: `simp` can prove it. +@[simp] theorem derivative_sub {f g : R[X]} : derivative (f - g) = derivative f - derivative g := LinearMap.map_sub derivative f g diff --git a/Mathlib/Algebra/Polynomial/Laurent.lean b/Mathlib/Algebra/Polynomial/Laurent.lean index c231267f42556..7210a0ec6104f 100644 --- a/Mathlib/Algebra/Polynomial/Laurent.lean +++ b/Mathlib/Algebra/Polynomial/Laurent.lean @@ -205,16 +205,16 @@ theorem _root_.Polynomial.toLaurent_X : (toLaurent Polynomial.X : R[T;T⁻¹]) = have : (Polynomial.X : R[X]) = monomial 1 1 := by simp [← C_mul_X_pow_eq_monomial] simp [this, Polynomial.toLaurent_C_mul_T] --- @[simp] -- Porting note (#10618): simp can prove this +@[simp] theorem _root_.Polynomial.toLaurent_one : (Polynomial.toLaurent : R[X] → R[T;T⁻¹]) 1 = 1 := map_one Polynomial.toLaurent --- @[simp] -- Porting note (#10618): simp can prove this +@[simp] theorem _root_.Polynomial.toLaurent_C_mul_eq (r : R) (f : R[X]) : toLaurent (Polynomial.C r * f) = C r * toLaurent f := by simp only [_root_.map_mul, Polynomial.toLaurent_C] --- @[simp] -- Porting note (#10618): simp can prove this +@[simp] theorem _root_.Polynomial.toLaurent_X_pow (n : ℕ) : toLaurent (X ^ n : R[X]) = T n := by simp only [map_pow, Polynomial.toLaurent_X, T_pow, mul_one] diff --git a/Mathlib/Algebra/Ring/Equiv.lean b/Mathlib/Algebra/Ring/Equiv.lean index 51c833ecb1d78..8518ce1badcb4 100644 --- a/Mathlib/Algebra/Ring/Equiv.lean +++ b/Mathlib/Algebra/Ring/Equiv.lean @@ -605,7 +605,7 @@ section Ring variable [NonAssocRing R] [NonAssocRing S] (f : R ≃+* S) --- Porting note (#10618): `simp` can now prove that, so we remove the `@[simp]` tag +@[simp] theorem map_neg_one : f (-1) = -1 := f.map_one ▸ f.map_neg 1 diff --git a/Mathlib/Algebra/Ring/Hom/Defs.lean b/Mathlib/Algebra/Ring/Hom/Defs.lean index 85b96d8538cfb..2ba354beb18de 100644 --- a/Mathlib/Algebra/Ring/Hom/Defs.lean +++ b/Mathlib/Algebra/Ring/Hom/Defs.lean @@ -649,8 +649,7 @@ theorem coe_fn_mkRingHomOfMulSelfOfTwoNeZero (h h_two h_one) : (f.mkRingHomOfMulSelfOfTwoNeZero h h_two h_one : β → α) = f := rfl --- Porting note (#10618): `simp` can prove this --- @[simp] +@[simp] theorem coe_addMonoidHom_mkRingHomOfMulSelfOfTwoNeZero (h h_two h_one) : (f.mkRingHomOfMulSelfOfTwoNeZero h h_two h_one : β →+ α) = f := by ext diff --git a/Mathlib/Algebra/Ring/Int.lean b/Mathlib/Algebra/Ring/Int.lean index 8fb0dd29430dc..45086c465483d 100644 --- a/Mathlib/Algebra/Ring/Int.lean +++ b/Mathlib/Algebra/Ring/Int.lean @@ -166,15 +166,13 @@ lemma even_mul_succ_self (n : ℤ) : Even (n * (n + 1)) := by lemma even_mul_pred_self (n : ℤ) : Even (n * (n - 1)) := by simpa [even_mul, parity_simps] using n.even_or_odd --- Porting note (#10618): was simp. simp can prove this. -@[norm_cast] lemma odd_coe_nat (n : ℕ) : Odd (n : ℤ) ↔ Odd n := by +@[simp, norm_cast] lemma odd_coe_nat (n : ℕ) : Odd (n : ℤ) ↔ Odd n := by rw [← not_even_iff_odd, ← Nat.not_even_iff_odd, even_coe_nat] @[simp] lemma natAbs_even : Even n.natAbs ↔ Even n := by simp [even_iff_two_dvd, dvd_natAbs, natCast_dvd.symm] --- Porting note (#10618): was simp. simp can prove this. ---@[simp] +@[simp] lemma natAbs_odd : Odd n.natAbs ↔ Odd n := by rw [← not_even_iff_odd, ← Nat.not_even_iff_odd, natAbs_even] diff --git a/Mathlib/Algebra/Ring/Subring/Basic.lean b/Mathlib/Algebra/Ring/Subring/Basic.lean index 781bbdda04c35..e16be89dd2957 100644 --- a/Mathlib/Algebra/Ring/Subring/Basic.lean +++ b/Mathlib/Algebra/Ring/Subring/Basic.lean @@ -387,11 +387,11 @@ def subtype (s : Subring R) : s →+* R := theorem coeSubtype : ⇑s.subtype = ((↑) : s → R) := rfl -@[norm_cast] -- Porting note (#10618): simp can prove this (removed `@[simp]`) +@[norm_cast] theorem coe_natCast : ∀ n : ℕ, ((n : s) : R) = n := map_natCast s.subtype -@[norm_cast] -- Porting note (#10618): simp can prove this (removed `@[simp]`) +@[norm_cast] theorem coe_intCast : ∀ n : ℤ, ((n : s) : R) = n := map_intCast s.subtype diff --git a/Mathlib/Algebra/Star/Subalgebra.lean b/Mathlib/Algebra/Star/Subalgebra.lean index a88d2390accc9..b2c1cc0ea0538 100644 --- a/Mathlib/Algebra/Star/Subalgebra.lean +++ b/Mathlib/Algebra/Star/Subalgebra.lean @@ -73,7 +73,6 @@ instance algebra (s : StarSubalgebra R A) : Algebra R s := instance starModule (s : StarSubalgebra R A) : StarModule R s where star_smul r a := Subtype.ext (star_smul r (a : A)) -@[simp, nolint simpNF] -- porting note (#10618): `simpNF` says `simp` can prove this, but it can't theorem mem_carrier {s : StarSubalgebra R A} {x : A} : x ∈ s.carrier ↔ x ∈ s := Iff.rfl diff --git a/Mathlib/FieldTheory/IntermediateField/Basic.lean b/Mathlib/FieldTheory/IntermediateField/Basic.lean index bd1efcb1176bb..23cb5fa01ffb4 100644 --- a/Mathlib/FieldTheory/IntermediateField/Basic.lean +++ b/Mathlib/FieldTheory/IntermediateField/Basic.lean @@ -288,7 +288,7 @@ namespace IntermediateField instance toField : Field S := S.toSubfield.toField -@[simp, norm_cast] +@[norm_cast] theorem coe_sum {ι : Type*} [Fintype ι] (f : ι → S) : (↑(∑ i, f i) : L) = ∑ i, (f i : L) := by classical induction' (Finset.univ : Finset ι) using Finset.induction_on with i s hi H From 10eb6ccf701b762128f9718ca0cc927a7b2c0bb0 Mon Sep 17 00:00:00 2001 From: damiano Date: Wed, 23 Oct 2024 02:43:28 +0000 Subject: [PATCH 311/505] feat(CI): add summary of benchmarking output (#16488) Add a GitHub action triggered on the `!bench` comments. Reports a summary of the changes by at least 10^9 instructions. [This PR](https://github.com/bryangingechen/mathlib4/pull/12#issuecomment-2328756897) contains various examples of reports generated by this script. [Zulip](https://leanprover.zulipchat.com/#narrow/stream/113488-general/topic/bench.20summaries) Co-authored-by: Michael Rothgang --- .github/workflows/bench_summary_comment.yml | 35 ++++ scripts/bench_summary.lean | 221 ++++++++++++++++++++ 2 files changed, 256 insertions(+) create mode 100644 .github/workflows/bench_summary_comment.yml create mode 100644 scripts/bench_summary.lean diff --git a/.github/workflows/bench_summary_comment.yml b/.github/workflows/bench_summary_comment.yml new file mode 100644 index 0000000000000..6672fb8443e98 --- /dev/null +++ b/.github/workflows/bench_summary_comment.yml @@ -0,0 +1,35 @@ +name: Bench output summary + +on: + issue_comment: + types: created + +jobs: + Produce_bench_summary: + name: Post summary of benchmarking results + if: github.event.issue.pull_request && (startsWith(github.event.comment.body, 'Here are the [benchmark results]')) + runs-on: ubuntu-latest + steps: + - name: install elan + run: | + set -o pipefail + curl -sSfL https://github.com/leanprover/elan/releases/download/v3.1.1/elan-x86_64-unknown-linux-gnu.tar.gz | tar xz + ./elan-init -y --default-toolchain none + echo "$HOME/.elan/bin" >> "${GITHUB_PATH}" + + - uses: actions/checkout@v4 + with: + ref: master + sparse-checkout: | + scripts/bench_summary.lean + + - name: Summarize bench output + run: | + { + cat scripts/bench_summary.lean + printf $'run_cmd BenchAction.addBenchSummaryComment %s "leanprover-community/mathlib4"' "${PR}" + } | + lake env lean --stdin + env: + PR: ${{ github.event.issue.number }} + GH_TOKEN: ${{secrets.GITHUB_TOKEN}} diff --git a/scripts/bench_summary.lean b/scripts/bench_summary.lean new file mode 100644 index 0000000000000..b7c521875b401 --- /dev/null +++ b/scripts/bench_summary.lean @@ -0,0 +1,221 @@ +/- +Copyright (c) 2024 Damiano Testa. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Damiano Testa +-/ + +import Lean.Elab.Command + +/-! +# Summary of `!bench` results + +This file contains a script that converts data retrieved from the speed-center into a +shorter, more accessible format, and post a comment with this summary on github. +-/ + +namespace BenchAction + +open Lean + +/-- `Bench` is a structure with the data used to compute the `!bench` summary. +It contains +* a string `file` (that could be `build`, `lint` or `~Mathlib.Path.To.File`); +* an integer `diff` representing the change in number of instructions for `file`; +* a float `reldiff` representing the percentage change in number of instructions for `file`. +-/ +structure Bench := + file : String + diff : Int + reldiff : Float + deriving FromJson, ToJson, Inhabited + +/-- `intDecs z exp prec` is a "generic" formatting of an integer `z`. +It writes `z` in the form `x.y * 10 ^ exp` (for non-negative integers `x`, `y` and `z`), +such that `y` has `prec` digits and returns +* the sign of `z` as a string (in fact, just either `+` or `-`); +* the integer `x`; +* the natural number `y` (that has `prec` digits). +-/ +def intDecs (z : Int) (exp : Nat := 9) (prec : Nat := 3) : String × Int × Nat := + let sgn := z.sign + let z := sgn * z + let p10 : Int := 10 ^ (exp - prec) + let idec := z / p10 + (if sgn < 0 then "-" else "+", idec / (10 ^ prec), (idec % 10 ^ prec).toNat) + +/-- `formatDiff z` uses `intDecs` to format an integer `z` as `±x.y⬝10⁹`. -/ +def formatDiff (z : Int) : String := + let (sgn, intDigs, decDigs) := intDecs z + s!"{sgn}{intDigs}.{decDigs}⬝10⁹" + +/-- Convert a `Float` into a formatted string of the form `±z.w%`. -/ +def formatPercent (reldiff : Float) : String := + -- shift by `2` twice: once to get a percentage, again for two decimal digits of precision + let reldiff := reldiff * 10 ^ 4 + let sgn : Int := if reldiff < 0 then -1 else 1 + let reldiff := (.ofInt sgn) * reldiff + let (sgn, intDigs, decDigs) := intDecs (sgn * reldiff.toUInt32.val) 0 2 + s!"({sgn}{intDigs}.{decDigs}%)" + +/-- +info: [(+0.0%), (+14.28%), (+0.20%), (-0.60%)] +--- +info: [+0.0⬝10⁹, +1.0⬝10⁹, +30.200⬝10⁹, -0.460⬝10⁹] +-/ +#guard_msgs in +run_cmd + let floats : Array Float := #[0, 1/7, 0.002, -0.006] + logInfo m!"{floats.map formatPercent}" + let ints : Array Int := #[0, 10^9, 302*10^8, -460000000] + logInfo m!"{ints.map formatDiff}" + +/-- +`formatFile file` converts a `String` into a formatted string of the form `` `file` ``, +removing leading non-letters. It is expected that `~` is the only leading non-letter. +-/ +def formatFile (file : String) : String := s!"`{file.dropWhile (!·.isAlpha)}`" + +/-- +`summary bc` converts a `Bench` into a formatted string of the form +``| `file` | ±x.y⬝10⁹ | ±z.w% |`` (technically, without the spaces). +-/ +def summary (bc : Bench) : String := + let reldiff := bc.reldiff + let (sgnf : Float) := if reldiff < 0 then -1 else 1 + let middle := [formatFile bc.file, formatDiff bc.diff, formatPercent (sgnf * reldiff)] + "|".intercalate (""::middle ++ [""]) + +/-- +`toTable bcs` formats an array of `Bench`es into a markdown table whose columns are +the file name, the absolute change in instruction counts and the relative change as a percentage. +A typical entry may look like ``|`Mathlib.Analysis.Seminorm`|+2.509⬝10⁹|(+1.41%)|``. +-/ +def toTable (bcs : Array Bench) : String := + let header := "|File|Instructions|%|\n|-|-:|:-:|" + "\n".intercalate (header :: (bcs.map summary).toList) + +/-- +`toCollapsibleTable bcs roundDiff` is similar to `toTable bcs`, except that it returns +output enclosed in a `
` html-block. +The `` part tallies the number of entries in `bcs` whose instructions increased +resp. decreased by at least the amount `roundDiff`. +-/ +def toCollapsibleTable (bcs : Array Bench) (roundDiff : Int) : String := + s!"
{bcs.size} files, Instructions {formatDiff <| roundDiff * 10 ^ 9}\ + \n\n{toTable (bcs.qsort (·.diff > ·.diff))}\n
\n" + +/-- Assuming that the input is a `json`-string formatted to produce an array of `Bench`, +`benchOutput` returns the "significant" changes in numbers of instructions as a string. -/ +def benchOutput (jsonInput : String) : IO String := do + let data ← IO.ofExcept (Json.parse jsonInput >>= fromJson? (α := Array Bench)) + -- `head` contains the changes for `build` and `lint`, + -- `data` contains the instruction changes for individual files: + -- each filename is prefixed by `~`. + let (head, data) := data.partition (·.file.take 1 != "~") + -- Partition the `Bench`es into "bins", i.e. the subsets of all `Bench`es whose difference + -- in instructions lies in an interval `[n·10⁹, (n+1)·10⁹)`. + -- The values `n` need not be consecutive: we only retain non-empty bins. + let grouped := ((data.groupByKey (·.diff / (10 ^ 9))).toArray.qsort (·.1 > ·.1)).toList + -- We consider `build` and `lint` as their own groups, in this order. + let sortHead := (head.qsort (·.file < ·.file)).toList.map (0, #[·]) + let togetherSorted := sortHead ++ grouped + -- For better formatting, we merge consecutive bins with just a single *file* into one. + -- This covers the two steps `ts1` and `ts2`. + -- For example, `[..., (bound₁, #[a₁]), (bound₂, #[a₂]), (bound₃, #[a₃]), ...]` gets collapsed to + -- `[..., (none, #[a₁, a₂, a₃]), ...]`. + -- The `boundᵢ` entry becomes `none` for the collapsed entries, so that we know that these + -- should be printed individually instead of inside a `
`-block. + -- A single bin with just a single file is also marked with `none`, for the same reason. + let ts1 := togetherSorted.groupBy (·.2.size == 1 && ·.2.size == 1) + let ts2 := List.join <| ts1.map fun l ↦ + if (l.getD 0 default).2.size == 1 then + [(none, l.foldl (· ++ ·.2) #[])] + else + l.map fun (n, ar) ↦ (some n, ar) + + let mut overall := [] + for (bound, gs) in ts2 do + overall := overall ++ [ + match bound with + -- These entries are from "singleton" files in their range; we print them individually. + | none => toTable gs + -- These get a collapsible summary instead. + | some roundedDiff => toCollapsibleTable gs roundedDiff] + return "\n".intercalate overall + +open Lean Elab Command in +/-- `addBenchSummaryComment PR repo tempFile` adds a summary of benchmarking results +as a comment to a pull request. It takes as input +* the number `PR` and the name `repo` as a `String` containing the relevant pull-request + (it reads and posts comments there) +* the `String` `tempFile` of a temporary file where the command stores transient information. + +The code itself interfaces with the shell to retrieve and process json data and eventually +uses `benchOutput`. +Here is a summary of the steps: +* retrieve the last comment to the PR (using `gh pr view ...`), +* check if it was posted by `leanprover-bot`, +* try to retrieve the source and target commits from the url that the bot posts + and store them in `source` and `target`, +* query the speed center for the benchmarking data (using `curl url`), +* format and filter the returned JSON data (various times), + saving intermediate steps into `tempFile` (using `jq` multiple times), +* process the final string to produce a summary (using `benchOutput`), +* finally post the resulting output to the PR (using `gh pr comment ...`). +-/ +def addBenchSummaryComment (PR : Nat) (repo : String) (tempFile : String := "benchOutput.json") : + CommandElabM Unit := do + let PR := s!"{PR}" + let jq := ".comments | last | select(.author.login==\"leanprover-bot\") | .body" + let gh_pr_comments : IO.Process.SpawnArgs := + { cmd := "gh", args := #["pr", "view", PR, "--repo", repo, "--json", "comments", "--jq", jq] } + -- This is the content of the last comment made by `leanprover-bot` to the PR `PR`. + let output ← IO.Process.run gh_pr_comments + -- URLs of benchmarking results have the form {something}/compare/source_sha/to/target_sha, + -- where source_sha and target_sha are the commit hashes of the revisions being benchmarked. + -- The comment contains such a URL (and only one); parse the revisions from the comment. + let frags := output.split (· == '/') + let some compIdx := frags.findIdx? (· == "compare") | + logInfo "No 'compare' found in URL." + return + let source := frags.getD (compIdx + 1) "" + let target := (frags.getD (compIdx + 3) "").takeWhile (· != ')') + if (source.length, target.length) != (36, 36) then + logInfo m!"Found\nsource: '{source}'\ntarget: '{target}'\ninstead of two commit hashes." + return + dbg_trace s!"Using commits\nsource: '{source}'\ntarget: '{target}'\n" + let curlSpeedCenter : IO.Process.SpawnArgs := + { cmd := "curl" + args := #[s!"http://speed.lean-fro.org/mathlib4/api/compare/{source}/to/{target}?all_values=true"] } + let bench ← IO.Process.run curlSpeedCenter + IO.FS.writeFile tempFile bench + -- Extract all instruction changes whose magnitude is larger than `threshold`. + let threshold := s!"{10 ^ 9}" + let jq1 : IO.Process.SpawnArgs := + { cmd := "jq" + args := #["-r", "--arg", "thr", threshold, + ".differences | .[] | ($thr|tonumber) as $th | + select(.dimension.metric == \"instructions\" and ((.diff >= $th) or (.diff <= -$th)))", + tempFile] } + let firstFilter ← IO.Process.run jq1 + IO.FS.writeFile tempFile firstFilter + -- Write these in compact form, in the format suitable for `benchOutput`. + let jq2 : IO.Process.SpawnArgs := + { cmd := "jq" + args := #["-c", "[{file: .dimension.benchmark, diff: .diff, reldiff: .reldiff}]", tempFile] } + let secondFilter ← IO.Process.run jq2 + IO.FS.writeFile tempFile secondFilter + let jq3 : IO.Process.SpawnArgs := + { cmd := "jq", args := #["-n", "reduce inputs as $in (null; . + $in)", tempFile] } + let thirdFilter ← IO.Process.run jq3 + let report ← benchOutput thirdFilter + IO.println report + -- Post the computed summary as a github comment. + let add_comment : IO.Process.SpawnArgs := + { cmd := "gh", args := #["pr", "comment", PR, "--repo", repo, "--body", report] } + let _ ← IO.Process.run add_comment + +end BenchAction + +-- CI adds the following line, replacing `putPR` with the PR number: +--run_cmd BenchAction.addBenchSummaryComment putPR "leanprover-community/mathlib4" From c48f8a2c507e553f98f2daadf5a9e63ccfe09047 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Wed, 23 Oct 2024 02:43:29 +0000 Subject: [PATCH 312/505] feat(CategoryTheory/Adjunction): the left partial adjoint (#17388) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Given a functor `F : D ⥤ C`, we define a functor`F.partialLeftAdjoint : F.PartialLeftAdjointSource ⥤ D` which is defined on a certain full subcategory of `C`. It satisfies similar properties to the left adjoint of `F` (if this existed). We show that the domain of definition of this partial left adjoint is stable under certain colimits. This shall be used in #17366 in order to show the existence of the pullback functor on categories of presheaves of modules. --- Mathlib.lean | 1 + Mathlib/CategoryTheory/Adjunction/Basic.lean | 14 ++ .../Adjunction/PartialAdjoint.lean | 182 ++++++++++++++++++ 3 files changed, 197 insertions(+) create mode 100644 Mathlib/CategoryTheory/Adjunction/PartialAdjoint.lean diff --git a/Mathlib.lean b/Mathlib.lean index 408df6cd2784b..a46fe860dbad2 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1454,6 +1454,7 @@ import Mathlib.CategoryTheory.Adjunction.Limits import Mathlib.CategoryTheory.Adjunction.Mates import Mathlib.CategoryTheory.Adjunction.Opposites import Mathlib.CategoryTheory.Adjunction.Over +import Mathlib.CategoryTheory.Adjunction.PartialAdjoint import Mathlib.CategoryTheory.Adjunction.Reflective import Mathlib.CategoryTheory.Adjunction.Restrict import Mathlib.CategoryTheory.Adjunction.Triple diff --git a/Mathlib/CategoryTheory/Adjunction/Basic.lean b/Mathlib/CategoryTheory/Adjunction/Basic.lean index d7c1293793302..634009df4e791 100644 --- a/Mathlib/CategoryTheory/Adjunction/Basic.lean +++ b/Mathlib/CategoryTheory/Adjunction/Basic.lean @@ -277,6 +277,20 @@ theorem eq_homEquiv_apply {A : C} {B : D} (f : F.obj A ⟶ B) (g : A ⟶ G.obj B g = adj.homEquiv A B f ↔ (adj.homEquiv A B).symm g = f := eq_unit_comp_map_iff adj f g +/-- If `adj : F ⊣ G`, and `X : C`, then `F.obj X` corepresents `Y ↦ (X ⟶ G.obj Y)`-/ +@[simps] +def corepresentableBy (X : C) : + (G ⋙ coyoneda.obj (Opposite.op X)).CorepresentableBy (F.obj X) where + homEquiv := adj.homEquiv _ _ + homEquiv_comp := by aesop_cat + +/-- If `adj : F ⊣ G`, and `Y : D`, then `G.obj Y` represents `X ↦ (F.obj X ⟶ Y)`-/ +@[simps] +def representableBy (Y : D) : + (F.op ⋙ yoneda.obj Y).RepresentableBy (G.obj Y) where + homEquiv := (adj.homEquiv _ _).symm + homEquiv_comp := by aesop_cat + end end Adjunction diff --git a/Mathlib/CategoryTheory/Adjunction/PartialAdjoint.lean b/Mathlib/CategoryTheory/Adjunction/PartialAdjoint.lean new file mode 100644 index 0000000000000..80ce40cc477e2 --- /dev/null +++ b/Mathlib/CategoryTheory/Adjunction/PartialAdjoint.lean @@ -0,0 +1,182 @@ +/- +Copyright (c) 2024 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ + +import Mathlib.CategoryTheory.Adjunction.Basic +import Mathlib.CategoryTheory.Limits.HasLimits +import Mathlib.CategoryTheory.Yoneda +import Mathlib.Order.CompleteLattice + +/-! +# Domain of definition of the partial left adjoint + +Given a functor `F : D ⥤ C`, we define a functor +`F.partialLeftAdjoint : F.PartialLeftAdjointSource ⥤ D` which is +defined on the full subcategory of `C` consisting of those objects `X : C` +such that `F ⋙ coyoneda.obj (op X) : D ⥤ Type _` is corepresentable. +We have a natural bijection +`(F.partialLeftAdjoint.obj X ⟶ Y) ≃ (X.obj ⟶ F.obj Y)` +that is similar to what we would expect for the image of the object `X` +by the left adjoint of `F`, if such an adjoint existed. + +Indeed, if the predicate `F.LeftAdjointObjIsDefined` which defines +the `F.PartialLeftAdjointSource` holds for all +objects `X : C`, then `F` has a left adjoint. + +When colimits indexed by a category `J` exist in `D`, we show that +the predicate `F.LeftAdjointObjIsDefined` is stable under colimits indexed by `J`. + +## TODO +* consider dualizing the results to right adjoints + +-/ + +universe v₁ v₂ u₁ u₂ + +namespace CategoryTheory + +namespace Functor + +open Category Opposite Limits + +variable {C : Type u₁} [Category.{v₁} C] {D : Type u₂} [Category.{v₂} D] (F : D ⥤ C) + +/-- Given a functor `F : D ⥤ C`, this is a predicate on objects `X : C` corresponding +to the domain of definition of the (partial) left adjoint of `F`. -/ +def LeftAdjointObjIsDefined (X : C) : Prop := IsCorepresentable (F ⋙ coyoneda.obj (op X)) + +lemma leftAdjointObjIsDefined_iff (X : C) : + F.LeftAdjointObjIsDefined X ↔ IsCorepresentable (F ⋙ coyoneda.obj (op X)) := by rfl + +variable {F} in +lemma leftAdjointObjIsDefined_of_adjunction {G : C ⥤ D} (adj : G ⊣ F) (X : C) : + F.LeftAdjointObjIsDefined X := + (adj.corepresentableBy X).isCorepresentable + +/-- The full subcategory where `F.partialLeftAdjoint` shall be defined. -/ +abbrev PartialLeftAdjointSource := FullSubcategory F.LeftAdjointObjIsDefined + +instance (X : F.PartialLeftAdjointSource) : + IsCorepresentable (F ⋙ coyoneda.obj (op X.obj)) := X.property + +/-- Given `F : D ⥤ C`, this is `F.partialLeftAdjoint` on objects: it sends +`X : C` such that `F.LeftAdjointObjIsDefined X` holds to an object of `D` +which represents the functor `F ⋙ coyoneda.obj (op X.obj)`. -/ +noncomputable def partialLeftAdjointObj (X : F.PartialLeftAdjointSource) : D := + (F ⋙ coyoneda.obj (op X.obj)).coreprX + +/-- Given `F : D ⥤ C`, this is the canonical bijection +`(F.partialLeftAdjointObj X ⟶ Y) ≃ (X.obj ⟶ F.obj Y)` +for all `X : F.PartialLeftAdjointSource` and `Y : D`. -/ +noncomputable def partialLeftAdjointHomEquiv {X : F.PartialLeftAdjointSource} {Y : D} : + (F.partialLeftAdjointObj X ⟶ Y) ≃ (X.obj ⟶ F.obj Y) := + (F ⋙ coyoneda.obj (op X.obj)).corepresentableBy.homEquiv + +lemma partialLeftAdjointHomEquiv_comp {X : F.PartialLeftAdjointSource} {Y Y' : D} + (f : F.partialLeftAdjointObj X ⟶ Y) (g : Y ⟶ Y') : + F.partialLeftAdjointHomEquiv (f ≫ g) = + F.partialLeftAdjointHomEquiv f ≫ F.map g := by + apply CorepresentableBy.homEquiv_comp + +/-- Given `F : D ⥤ C`, this is `F.partialLeftAdjoint` on morphisms. -/ +noncomputable def partialLeftAdjointMap {X Y : F.PartialLeftAdjointSource} + (f : X ⟶ Y) : F.partialLeftAdjointObj X ⟶ F.partialLeftAdjointObj Y := + F.partialLeftAdjointHomEquiv.symm (f ≫ F.partialLeftAdjointHomEquiv (𝟙 _)) + +@[simp] +lemma partialLeftAdjointHomEquiv_map {X Y : F.PartialLeftAdjointSource} + (f : X ⟶ Y) : + F.partialLeftAdjointHomEquiv (F.partialLeftAdjointMap f) = + by exact f ≫ F.partialLeftAdjointHomEquiv (𝟙 _) := by + simp [partialLeftAdjointMap] + +lemma partialLeftAdjointHomEquiv_map_comp {X X' : F.PartialLeftAdjointSource} {Y : D} + (f : X ⟶ X') (g : F.partialLeftAdjointObj X' ⟶ Y) : + F.partialLeftAdjointHomEquiv (F.partialLeftAdjointMap f ≫ g) = + by exact f ≫ F.partialLeftAdjointHomEquiv g := by + rw [partialLeftAdjointHomEquiv_comp, partialLeftAdjointHomEquiv_map, assoc, + ← partialLeftAdjointHomEquiv_comp, id_comp] + +/-- Given `F : D ⥤ C`, this is the partial adjoint functor `F.PartialLeftAdjointSource ⥤ D`. -/ +@[simps] +noncomputable def partialLeftAdjoint : F.PartialLeftAdjointSource ⥤ D where + obj := F.partialLeftAdjointObj + map := F.partialLeftAdjointMap + map_id X := by + apply F.partialLeftAdjointHomEquiv.injective + dsimp + rw [partialLeftAdjointHomEquiv_map] + erw [id_comp] + map_comp {X Y Z} f g := by + apply F.partialLeftAdjointHomEquiv.injective + dsimp + rw [partialLeftAdjointHomEquiv_map, partialLeftAdjointHomEquiv_comp, + partialLeftAdjointHomEquiv_map, assoc] + erw [assoc] + rw [← F.partialLeftAdjointHomEquiv_comp, id_comp, + partialLeftAdjointHomEquiv_map] + +variable {F} + +lemma isRightAdjoint_of_leftAdjointObjIsDefined_eq_top + (h : F.LeftAdjointObjIsDefined = ⊤) : F.IsRightAdjoint := by + replace h : ∀ X, IsCorepresentable (F ⋙ coyoneda.obj (op X)) := fun X ↦ by + simp only [← leftAdjointObjIsDefined_iff, h, Pi.top_apply, Prop.top_eq_true] + exact (Adjunction.adjunctionOfEquivLeft + (fun X Y ↦ (F ⋙ coyoneda.obj (op X)).corepresentableBy.homEquiv) + (fun X Y Y' g f ↦ by apply CorepresentableBy.homEquiv_comp)).isRightAdjoint + +variable (F) in +lemma isRightAdjoint_iff_leftAdjointObjIsDefined_eq_top : + F.IsRightAdjoint ↔ F.LeftAdjointObjIsDefined = ⊤ := by + refine ⟨fun h ↦ ?_, isRightAdjoint_of_leftAdjointObjIsDefined_eq_top⟩ + ext X + simpa only [Pi.top_apply, Prop.top_eq_true, iff_true] + using leftAdjointObjIsDefined_of_adjunction (Adjunction.ofIsRightAdjoint F) X + +/-- Auxiliary definition for `leftAdjointObjIsDefined_of_isColimit`. -/ +noncomputable def corepresentableByCompCoyonedaObjOfIsColimit {J : Type*} [Category J] + {R : J ⥤ F.PartialLeftAdjointSource} + {c : Cocone (R ⋙ fullSubcategoryInclusion _)} (hc : IsColimit c) + {c' : Cocone (R ⋙ F.partialLeftAdjoint)} (hc' : IsColimit c') : + (F ⋙ coyoneda.obj (op c.pt)).CorepresentableBy c'.pt where + homEquiv {Y} := + { toFun := fun f ↦ hc.desc (Cocone.mk _ + { app := fun j ↦ F.partialLeftAdjointHomEquiv (c'.ι.app j ≫ f) + naturality := fun j j' φ ↦ by + dsimp + rw [comp_id, ← c'.w φ, ← partialLeftAdjointHomEquiv_map_comp, assoc] + dsimp }) + invFun := fun g ↦ hc'.desc (Cocone.mk _ + { app := fun j ↦ F.partialLeftAdjointHomEquiv.symm (c.ι.app j ≫ g) + naturality := fun j j' φ ↦ by + apply F.partialLeftAdjointHomEquiv.injective + have := c.w φ + dsimp at this ⊢ + rw [comp_id, Equiv.apply_symm_apply, partialLeftAdjointHomEquiv_map_comp, + Equiv.apply_symm_apply, reassoc_of% this] }) + left_inv := fun f ↦ hc'.hom_ext (fun j ↦ by simp) + right_inv := fun g ↦ hc.hom_ext (fun j ↦ by simp) } + homEquiv_comp {Y Y'} g f := hc.hom_ext (fun j ↦ by + dsimp + simp only [IsColimit.fac, IsColimit.fac_assoc, partialLeftAdjointHomEquiv_comp, + F.map_comp, assoc] ) + +lemma leftAdjointObjIsDefined_of_isColimit {J : Type*} [Category J] {R : J ⥤ C} {c : Cocone R} + (hc : IsColimit c) [HasColimitsOfShape J D] + (h : ∀ (j : J), F.LeftAdjointObjIsDefined (R.obj j)) : + F.LeftAdjointObjIsDefined c.pt := + (corepresentableByCompCoyonedaObjOfIsColimit + (R := FullSubcategory.lift _ R h) hc (colimit.isColimit _)).isCorepresentable + +lemma leftAdjointObjIsDefined_colimit {J : Type*} [Category J] (R : J ⥤ C) + [HasColimit R] [HasColimitsOfShape J D] + (h : ∀ (j : J), F.LeftAdjointObjIsDefined (R.obj j)) : + F.LeftAdjointObjIsDefined (colimit R) := + leftAdjointObjIsDefined_of_isColimit (colimit.isColimit R) h + +end Functor + +end CategoryTheory From 1eba96a6aa80fec5478de840c90f232034571796 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 23 Oct 2024 02:43:31 +0000 Subject: [PATCH 313/505] chore: use newly introduced finset notation (#17974) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit For `s : Finset α`, replace `s.card` with `#s`, `s.filter p` with `{x ∈ s | p x}`, `∑ x ∈ s.filter p, f x` with `∑ x ∈ s with p x, f x`. Rewrap lines around and open the `Finset` locale where necessary. --- .../SimplicialSet/Basic.lean | 3 +- Mathlib/Analysis/BoxIntegral/Basic.lean | 14 +-- .../Analysis/BoxIntegral/Partition/Basic.lean | 4 +- .../BoxIntegral/Partition/Tagged.lean | 9 +- .../Analysis/Calculus/FDeriv/Analytic.lean | 2 +- Mathlib/Analysis/Convex/Birkhoff.lean | 14 +-- Mathlib/Analysis/Convex/Caratheodory.lean | 10 +- Mathlib/Analysis/Convex/Combination.lean | 5 +- Mathlib/Analysis/Convex/Jensen.lean | 4 +- Mathlib/Analysis/Convex/Radon.lean | 34 +++---- Mathlib/Analysis/MeanInequalities.lean | 6 +- .../NormedSpace/Multilinear/Basic.lean | 6 +- .../NormedSpace/Multilinear/Curry.lean | 12 +-- Mathlib/Analysis/SpecificLimits/FloorPow.lean | 17 ++-- .../Computability/AkraBazzi/AkraBazzi.lean | 20 ++-- Mathlib/FieldTheory/Finite/Basic.lean | 14 +-- Mathlib/Geometry/Euclidean/Circumcenter.lean | 8 +- Mathlib/Geometry/Euclidean/MongePoint.lean | 6 +- Mathlib/GroupTheory/Index.lean | 2 +- Mathlib/GroupTheory/Perm/Cycle/Basic.lean | 60 ++++++------ Mathlib/GroupTheory/Perm/Cycle/Factors.lean | 15 ++- Mathlib/GroupTheory/Perm/Support.lean | 37 ++++--- .../GroupTheory/SpecificGroups/Cyclic.lean | 46 ++++----- Mathlib/InformationTheory/Hamming.lean | 6 +- .../AffineSpace/Combination.lean | 50 +++++----- .../AffineSpace/FiniteDimensional.lean | 17 ++-- .../AffineSpace/Independent.lean | 14 +-- Mathlib/LinearAlgebra/Lagrange.lean | 57 +++++------ .../LinearAlgebra/Matrix/AbsoluteValue.lean | 6 +- .../Matrix/Determinant/Basic.lean | 12 +-- Mathlib/LinearAlgebra/Multilinear/Basic.lean | 58 +++++------ .../LinearAlgebra/QuadraticForm/Basic.lean | 2 +- Mathlib/Logic/Denumerable.lean | 10 +- .../MeasureTheory/Function/SimpleFunc.lean | 6 +- Mathlib/MeasureTheory/Integral/Bochner.lean | 4 +- Mathlib/MeasureTheory/Integral/SetToL1.lean | 17 ++-- Mathlib/NumberTheory/ArithmeticFunction.lean | 8 +- .../DiophantineApproximation.lean | 2 +- Mathlib/NumberTheory/Divisors.lean | 19 ++-- Mathlib/NumberTheory/EulerProduct/Basic.lean | 14 +-- Mathlib/NumberTheory/JacobiSum/Basic.lean | 2 +- .../LegendreSymbol/GaussEisensteinLemmas.lean | 71 +++++++------- .../LegendreSymbol/QuadraticChar/Basic.lean | 2 +- Mathlib/NumberTheory/MaricaSchoenheim.lean | 6 +- .../NumberTheory/NumberField/Embeddings.lean | 34 +++---- .../NumberField/EquivReindex.lean | 5 +- Mathlib/NumberTheory/PrimeCounting.lean | 8 +- Mathlib/NumberTheory/Primorial.lean | 7 +- Mathlib/NumberTheory/SiegelsLemma.lean | 12 +-- Mathlib/Order/Interval/Finset/Basic.lean | 91 +++++++---------- Mathlib/Order/Interval/Finset/Box.lean | 8 +- Mathlib/Order/Interval/Finset/Defs.lean | 98 +++++++++---------- Mathlib/Order/Interval/Finset/Fin.lean | 25 ++--- Mathlib/Order/Interval/Finset/Nat.lean | 25 ++--- Mathlib/Order/Partition/Equipartition.lean | 82 +++++++--------- Mathlib/Order/Partition/Finpartition.lean | 46 +++++---- Mathlib/Order/SupClosed.lean | 4 +- Mathlib/Probability/StrongLaw.lean | 3 +- Mathlib/RingTheory/GradedAlgebra/Radical.lean | 10 +- Mathlib/RingTheory/IntegralDomain.lean | 15 ++- .../RingTheory/MvPolynomial/Homogeneous.lean | 3 +- .../MvPolynomial/Symmetric/Defs.lean | 2 +- .../Symmetric/FundamentalTheorem.lean | 4 +- .../Symmetric/NewtonIdentities.lean | 65 ++++++------ .../MvPolynomial/WeightedHomogeneous.lean | 2 +- Mathlib/RingTheory/MvPowerSeries/Basic.lean | 6 +- Mathlib/RingTheory/Polynomial/Vieta.lean | 9 +- Mathlib/RingTheory/Prime.lean | 2 +- Mathlib/RingTheory/RootsOfUnity/Basic.lean | 4 +- Mathlib/RingTheory/RootsOfUnity/Lemmas.lean | 2 +- Mathlib/Tactic/Positivity/Finset.lean | 12 +-- .../Algebra/InfiniteSum/Constructions.lean | 8 +- .../Topology/Algebra/InfiniteSum/Group.lean | 4 +- .../Topology/Algebra/InfiniteSum/Order.lean | 4 +- Mathlib/Topology/Instances/ENNReal.lean | 6 +- test/antidiagonal.lean | 4 +- test/positivity.lean | 4 +- 77 files changed, 620 insertions(+), 725 deletions(-) diff --git a/Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean b/Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean index 32b22e7b91ce5..814de7455bce4 100644 --- a/Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean +++ b/Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean @@ -221,8 +221,7 @@ def const (n : ℕ) (i k : Fin (n+3)) (m : SimplexCategoryᵒᵖ) : Λ[n+2, i].o This edge only exists if `{i, a, b}` has cardinality less than `n`. -/ @[simps] -def edge (n : ℕ) (i a b : Fin (n+1)) (hab : a ≤ b) (H : Finset.card {i, a, b} ≤ n) : - Λ[n, i] _[1] := by +def edge (n : ℕ) (i a b : Fin (n+1)) (hab : a ≤ b) (H : #{i, a, b} ≤ n) : Λ[n, i] _[1] := by refine ⟨standardSimplex.edge n a b hab, ?range⟩ case range => suffices ∃ x, ¬i = x ∧ ¬a = x ∧ ¬b = x by diff --git a/Mathlib/Analysis/BoxIntegral/Basic.lean b/Mathlib/Analysis/BoxIntegral/Basic.lean index 77776a456f6e3..44c6028cdd79d 100644 --- a/Mathlib/Analysis/BoxIntegral/Basic.lean +++ b/Mathlib/Analysis/BoxIntegral/Basic.lean @@ -512,8 +512,8 @@ theorem dist_integralSum_sum_integral_le_of_memBaseSet_of_iUnion_eq (h : Integra -- Let us prove that the distance is less than or equal to `ε + δ` for all positive `δ`. refine le_of_forall_pos_le_add fun δ δ0 => ?_ -- First we choose some constants. - set δ' : ℝ := δ / (π₀.boxes.card + 1) - have H0 : 0 < (π₀.boxes.card + 1 : ℝ) := Nat.cast_add_one_pos _ + set δ' : ℝ := δ / (#π₀.boxes + 1) + have H0 : 0 < (#π₀.boxes + 1 : ℝ) := Nat.cast_add_one_pos _ have δ'0 : 0 < δ' := div_pos δ0 H0 set C := max π₀.distortion π₀.compl.distortion /- Next we choose a tagged partition of each `J ∈ π₀` such that the integral sum of `f` over this @@ -665,7 +665,7 @@ theorem integrable_of_bounded_and_ae_continuousWithinAt [CompleteSpace E] {I : B let t₂ (J : Box ι) : ℝⁿ := (π₂.infPrepartition π₁.toPrepartition).tag J let B := (π₁.toPrepartition ⊓ π₂.toPrepartition).boxes classical - let B' := B.filter (fun J ↦ J.toSet ⊆ U) + let B' := {J ∈ B | J.toSet ⊆ U} have hB' : B' ⊆ B := B.filter_subset (fun J ↦ J.toSet ⊆ U) have μJ_ne_top : ∀ J ∈ B, μ J ≠ ⊤ := fun J hJ ↦ lt_top_iff_ne_top.1 <| lt_of_le_of_lt (μ.mono (Prepartition.le_of_mem' _ J hJ)) μI @@ -795,7 +795,7 @@ theorem HasIntegral.of_bRiemann_eq_false_of_forall_isLittleO (hl : l.bRiemann = /- For the boxes such that `π.tag J ∈ s`, we use the fact that at most `2 ^ #ι` boxes have the same tag. -/ specialize hlH hsne - have : ∀ J ∈ π.boxes.filter fun J => π.tag J ∈ s, + have : ∀ J ∈ {J ∈ π.boxes | π.tag J ∈ s}, dist (vol J (f <| π.tag J)) (g J) ≤ εs (π.tag J) := fun J hJ ↦ by rw [Finset.mem_filter] at hJ; cases' hJ with hJ hJs refine Hδ₁ c _ ⟨π.tag_mem_Icc _, hJs⟩ _ (hεs0 _) _ (π.le_of_mem' _ hJ) ?_ @@ -815,9 +815,9 @@ theorem HasIntegral.of_bRiemann_eq_false_of_forall_isLittleO (hl : l.bRiemann = In this case the estimate is straightforward. -/ -- Porting note: avoided strange elaboration issues by rewriting using `calc` calc - dist (∑ J ∈ π.boxes.filter (¬tag π · ∈ s), vol J (f (tag π J))) - (∑ J ∈ π.boxes.filter (¬tag π · ∈ s), g J) - ≤ ∑ J ∈ π.boxes.filter (¬tag π · ∈ s), ε' * B J := dist_sum_sum_le_of_le _ fun J hJ ↦ by + dist (∑ J ∈ π.boxes with ¬tag π J ∈ s, vol J (f (tag π J))) + (∑ J ∈ π.boxes with ¬tag π J ∈ s, g J) + ≤ ∑ J ∈ π.boxes with ¬tag π J ∈ s, ε' * B J := dist_sum_sum_le_of_le _ fun J hJ ↦ by rw [Finset.mem_filter] at hJ; cases' hJ with hJ hJs refine Hδ₂ c _ ⟨π.tag_mem_Icc _, hJs⟩ _ ε'0 _ (π.le_of_mem' _ hJ) ?_ (fun hH => hπδ.2 hH J hJ) fun hD => (Finset.le_sup hJ).trans (hπδ.3 hD) diff --git a/Mathlib/Analysis/BoxIntegral/Partition/Basic.lean b/Mathlib/Analysis/BoxIntegral/Partition/Basic.lean index c8af644f555c9..1686d1eac27e6 100644 --- a/Mathlib/Analysis/BoxIntegral/Partition/Basic.lean +++ b/Mathlib/Analysis/BoxIntegral/Partition/Basic.lean @@ -177,7 +177,7 @@ theorem injOn_setOf_mem_Icc_setOf_lower_eq (x : ι → ℝ) : /-- The set of boxes of a prepartition that contain `x` in their closures has cardinality at most `2 ^ Fintype.card ι`. -/ theorem card_filter_mem_Icc_le [Fintype ι] (x : ι → ℝ) : - (π.boxes.filter fun J : Box ι => x ∈ Box.Icc J).card ≤ 2 ^ Fintype.card ι := by + #{J ∈ π.boxes | x ∈ Box.Icc J} ≤ 2 ^ Fintype.card ι := by rw [← Fintype.card_set] refine Finset.card_le_card_of_injOn (fun J : Box ι => { i | J.lower i = x i }) (fun _ _ => Finset.mem_univ _) ?_ @@ -517,7 +517,7 @@ instance : SemilatticeInf (Prepartition I) := /-- The prepartition with boxes `{J ∈ π | p J}`. -/ @[simps] def filter (π : Prepartition I) (p : Box ι → Prop) : Prepartition I where - boxes := π.boxes.filter p + boxes := {J ∈ π.boxes | p J} le_of_mem' _ hJ := π.le_of_mem (mem_filter.1 hJ).1 pairwiseDisjoint _ h₁ _ h₂ := π.disjoint_coe_of_mem (mem_filter.1 h₁).1 (mem_filter.1 h₂).1 diff --git a/Mathlib/Analysis/BoxIntegral/Partition/Tagged.lean b/Mathlib/Analysis/BoxIntegral/Partition/Tagged.lean index 967500dc84169..9710dd56108b3 100644 --- a/Mathlib/Analysis/BoxIntegral/Partition/Tagged.lean +++ b/Mathlib/Analysis/BoxIntegral/Partition/Tagged.lean @@ -25,10 +25,8 @@ rectangular box, box partition noncomputable section +open Finset Function ENNReal NNReal Set open scoped Classical -open ENNReal NNReal - -open Set Function namespace BoxIntegral @@ -204,10 +202,9 @@ theorem isHenstock_biUnionTagged {π : Prepartition I} {πi : ∀ J, TaggedPrepa /-- In a Henstock prepartition, there are at most `2 ^ Fintype.card ι` boxes with a given tag. -/ theorem IsHenstock.card_filter_tag_eq_le [Fintype ι] (h : π.IsHenstock) (x : ι → ℝ) : - (π.boxes.filter fun J => π.tag J = x).card ≤ 2 ^ Fintype.card ι := + #{J ∈ π.boxes | π.tag J = x} ≤ 2 ^ Fintype.card ι := calc - (π.boxes.filter fun J => π.tag J = x).card ≤ - (π.boxes.filter fun J : Box ι => x ∈ Box.Icc J).card := by + #{J ∈ π.boxes | π.tag J = x} ≤ #{J ∈ π.boxes | x ∈ Box.Icc J} := by refine Finset.card_le_card fun J hJ => ?_ rw [Finset.mem_filter] at hJ ⊢; rcases hJ with ⟨hJ, rfl⟩ exact ⟨hJ, h J hJ⟩ diff --git a/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean b/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean index 80735fc0538cf..4947d3460259a 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean @@ -619,7 +619,7 @@ theorem changeOrigin_toFormalMultilinearSeries [DecidableEq ι] : rw [card_compl, Fintype.card_fin, Finset.card_singleton, Nat.add_sub_cancel_left]⟩) · use fun _ _ ↦ (singleton_injective <| compl_injective <| Subtype.ext_iff.mp ·) intro ⟨s, hs⟩ - have h : sᶜ.card = 1 := by rw [card_compl, hs, Fintype.card_fin, Nat.add_sub_cancel] + have h : #sᶜ = 1 := by rw [card_compl, hs, Fintype.card_fin, Nat.add_sub_cancel] obtain ⟨a, ha⟩ := card_eq_one.mp h exact ⟨a, Subtype.ext (compl_eq_comm.mp ha)⟩ rw [Function.comp_apply, Subtype.coe_mk, compl_singleton, piecewise_erase_univ, diff --git a/Mathlib/Analysis/Convex/Birkhoff.lean b/Mathlib/Analysis/Convex/Birkhoff.lean index a754f5db135c9..702bc496c5a0d 100644 --- a/Mathlib/Analysis/Convex/Birkhoff.lean +++ b/Mathlib/Analysis/Convex/Birkhoff.lean @@ -46,13 +46,13 @@ private lemma exists_perm_eq_zero_implies_eq_zero [Nonempty n] {s : R} (hs : 0 < (hM : ∃ M' ∈ doublyStochastic R n, M = s • M') : ∃ σ : Equiv.Perm n, ∀ i j, M i j = 0 → σ.permMatrix R i j = 0 := by rw [exists_mem_doublyStochastic_eq_smul_iff hs.le] at hM - let f (i : n) : Finset n := univ.filter (M i · ≠ 0) - have hf (A : Finset n) : A.card ≤ (A.biUnion f).card := by + let f (i : n) : Finset n := {j | M i j ≠ 0} + have hf (A : Finset n) : #A ≤ #(A.biUnion f) := by have (i) : ∑ j ∈ f i, M i j = s := by simp [sum_subset (filter_subset _ _), hM.2.1] - have h₁ : ∑ i ∈ A, ∑ j ∈ f i, M i j = A.card * s := by simp [this] - have h₂ : ∑ i, ∑ j ∈ A.biUnion f, M i j = (A.biUnion f).card * s := by + have h₁ : ∑ i ∈ A, ∑ j ∈ f i, M i j = #A * s := by simp [this] + have h₂ : ∑ i, ∑ j ∈ A.biUnion f, M i j = #(A.biUnion f) * s := by simp [sum_comm (t := A.biUnion f), hM.2.2, mul_comm s] - suffices A.card * s ≤ (A.biUnion f).card * s by exact_mod_cast le_of_mul_le_mul_right this hs + suffices #A * s ≤ #(A.biUnion f) * s by exact_mod_cast le_of_mul_le_mul_right this hs rw [← h₁, ← h₂] trans ∑ i ∈ A, ∑ j ∈ A.biUnion f, M i j · refine sum_le_sum fun i hi => ?_ @@ -86,7 +86,7 @@ private lemma doublyStochastic_sum_perm_aux (M : Matrix n n R) ∃ w : Equiv.Perm n → R, (∀ σ, 0 ≤ w σ) ∧ ∑ σ, w σ • σ.permMatrix R = M := by rcases isEmpty_or_nonempty n case inl => exact ⟨1, by simp, Subsingleton.elim _ _⟩ - set d : ℕ := (Finset.univ.filter fun i : n × n => M i.1 i.2 ≠ 0).card with ← hd + set d : ℕ := #{i : n × n | M i.1 i.2 ≠ 0} with ← hd clear_value d induction d using Nat.strongRecOn generalizing M s case ind d ih => @@ -114,7 +114,7 @@ private lemma doublyStochastic_sum_perm_aux (M : Matrix n n R) split case isTrue h => exact (hi' i' (by simp)).trans_eq (by rw [h]) case isFalse h => exact hM.1 _ _ - have hd' : (univ.filter fun i : n × n => N i.1 i.2 ≠ 0).card < d := by + have hd' : #{i : n × n | N i.1 i.2 ≠ 0} < d := by rw [← hd] refine card_lt_card ?_ rw [ssubset_iff_of_subset (monotone_filter_right _ _)] diff --git a/Mathlib/Analysis/Convex/Caratheodory.lean b/Mathlib/Analysis/Convex/Caratheodory.lean index 356e64e002984..59b4f09c6a3ec 100644 --- a/Mathlib/Analysis/Convex/Caratheodory.lean +++ b/Mathlib/Analysis/Convex/Caratheodory.lean @@ -55,7 +55,7 @@ theorem mem_convexHull_erase [DecidableEq E] {t : Finset E} (h : ¬AffineIndepen obtain ⟨g, gcombo, gsum, gpos⟩ := exists_nontrivial_relation_sum_zero_of_not_affine_ind h replace gpos := exists_pos_of_sum_zero_of_exists_nonzero g gsum gpos clear h - let s := @Finset.filter _ (fun z => 0 < g z) (fun _ => LinearOrder.decidableLT _ _) t + let s := {z ∈ t | 0 < g z} obtain ⟨i₀, mem, w⟩ : ∃ i₀ ∈ s, ∀ i ∈ s, f i₀ / g i₀ ≤ f i / g i := by apply s.exists_min_image fun z => f z / g z obtain ⟨x, hx, hgx⟩ : ∃ x ∈ t, 0 < g x := gpos @@ -117,13 +117,13 @@ theorem minCardFinsetOfMemConvexHull_nonempty : (minCardFinsetOfMemConvexHull hx exact ⟨x, mem_minCardFinsetOfMemConvexHull hx⟩ theorem minCardFinsetOfMemConvexHull_card_le_card {t : Finset E} (ht₁ : ↑t ⊆ s) - (ht₂ : x ∈ convexHull 𝕜 (t : Set E)) : (minCardFinsetOfMemConvexHull hx).card ≤ t.card := + (ht₂ : x ∈ convexHull 𝕜 (t : Set E)) : #(minCardFinsetOfMemConvexHull hx) ≤ #t := Function.argminOn_le _ _ _ (by exact ⟨ht₁, ht₂⟩) theorem affineIndependent_minCardFinsetOfMemConvexHull : AffineIndependent 𝕜 ((↑) : minCardFinsetOfMemConvexHull hx → E) := by - let k := (minCardFinsetOfMemConvexHull hx).card - 1 - have hk : (minCardFinsetOfMemConvexHull hx).card = k + 1 := + let k := #(minCardFinsetOfMemConvexHull hx) - 1 + have hk : #(minCardFinsetOfMemConvexHull hx) = k + 1 := (Nat.succ_pred_eq_of_pos (Finset.card_pos.mpr (minCardFinsetOfMemConvexHull_nonempty hx))).symm classical by_contra h @@ -163,7 +163,7 @@ theorem eq_pos_convex_span_of_mem_convexHull {x : E} (hx : x ∈ convexHull 𝕜 obtain ⟨t, ht₁, ht₂, ht₃⟩ := hx simp only [t.convexHull_eq, exists_prop, Set.mem_setOf_eq] at ht₃ obtain ⟨w, hw₁, hw₂, hw₃⟩ := ht₃ - let t' := t.filter fun i => w i ≠ 0 + let t' := {i ∈ t | w i ≠ 0} refine ⟨t', t'.fintypeCoeSort, ((↑) : t' → E), w ∘ ((↑) : t' → E), ?_, ?_, ?_, ?_, ?_⟩ · rw [Subtype.range_coe_subtype] exact Subset.trans (Finset.filter_subset _ t) ht₁ diff --git a/Mathlib/Analysis/Convex/Combination.lean b/Mathlib/Analysis/Convex/Combination.lean index 87bc080685121..6fe599b4f51f7 100644 --- a/Mathlib/Analysis/Convex/Combination.lean +++ b/Mathlib/Analysis/Convex/Combination.lean @@ -124,8 +124,7 @@ theorem Finset.centerMass_subset {t' : Finset ι} (ht : t ⊆ t') (h : ∀ i ∈ intro i hit' hit rw [h i hit' hit, zero_smul, smul_zero] -theorem Finset.centerMass_filter_ne_zero : - (t.filter fun i => w i ≠ 0).centerMass w z = t.centerMass w z := +theorem Finset.centerMass_filter_ne_zero : {i ∈ t | w i ≠ 0}.centerMass w z = t.centerMass w z := Finset.centerMass_subset z (filter_subset _ _) fun i hit hit' => by simpa only [hit, mem_filter, true_and, Ne, Classical.not_not] using hit' @@ -264,7 +263,7 @@ theorem Finset.centroid_mem_convexHull (s : Finset E) (hs : s.Nonempty) : rw [s.centroid_eq_centerMass hs] apply s.centerMass_id_mem_convexHull · simp only [inv_nonneg, imp_true_iff, Nat.cast_nonneg, Finset.centroidWeights_apply] - · have hs_card : (s.card : R) ≠ 0 := by simp [Finset.nonempty_iff_ne_empty.mp hs] + · have hs_card : (#s : R) ≠ 0 := by simp [Finset.nonempty_iff_ne_empty.mp hs] simp only [hs_card, Finset.sum_const, nsmul_eq_mul, mul_inv_cancel₀, Ne, not_false_iff, Finset.centroidWeights_apply, zero_lt_one] diff --git a/Mathlib/Analysis/Convex/Jensen.lean b/Mathlib/Analysis/Convex/Jensen.lean index 93f195e014393..90692868f122b 100644 --- a/Mathlib/Analysis/Convex/Jensen.lean +++ b/Mathlib/Analysis/Convex/Jensen.lean @@ -271,8 +271,8 @@ lemma ConvexOn.exists_ge_of_centerMass {t : Finset ι} (h : ConvexOn 𝕜 s f) ∃ i ∈ t, f (t.centerMass w p) ≤ f (p i) := by set y := t.centerMass w p -- TODO: can `rsuffices` be used to write the `exact` first, then the proof of this obtain? - obtain ⟨i, hi, hfi⟩ : ∃ i ∈ t.filter fun i => w i ≠ 0, w i • f y ≤ w i • (f ∘ p) i := by - have hw' : (0 : 𝕜) < ∑ i ∈ filter (fun i => w i ≠ 0) t, w i := by rwa [sum_filter_ne_zero] + obtain ⟨i, hi, hfi⟩ : ∃ i ∈ {i ∈ t | w i ≠ 0}, w i • f y ≤ w i • (f ∘ p) i := by + have hw' : (0 : 𝕜) < ∑ i ∈ t with w i ≠ 0, w i := by rwa [sum_filter_ne_zero] refine exists_le_of_sum_le (nonempty_of_sum_ne_zero hw'.ne') ?_ rw [← sum_smul, ← smul_le_smul_iff_of_pos_left (inv_pos.2 hw'), inv_smul_smul₀ hw'.ne', ← centerMass, centerMass_filter_ne_zero] diff --git a/Mathlib/Analysis/Convex/Radon.lean b/Mathlib/Analysis/Convex/Radon.lean index f297649b6f470..2a65ed2e8986a 100644 --- a/Mathlib/Analysis/Convex/Radon.lean +++ b/Mathlib/Analysis/Convex/Radon.lean @@ -41,8 +41,8 @@ theorem radon_partition {f : ι → E} (h : ¬ AffineIndependent 𝕜 f) : rw [affineIndependent_iff] at h push_neg at h obtain ⟨s, w, h_wsum, h_vsum, nonzero_w_index, h1, h2⟩ := h - let I : Finset ι := s.filter fun i ↦ 0 ≤ w i - let J : Finset ι := s.filter fun i ↦ w i < 0 + let I : Finset ι := {i ∈ s | 0 ≤ w i} + let J : Finset ι := {i ∈ s | w i < 0} let p : E := centerMass I w f -- point of intersection have hJI : ∑ j ∈ J, w j + ∑ i ∈ I, w i = 0 := by simpa only [h_wsum, not_lt] using sum_filter_add_sum_filter_not s (fun i ↦ w i < 0) w @@ -66,8 +66,8 @@ open Module /-- Corner case for `helly_theorem'`. -/ private lemma helly_theorem_corner {F : ι → Set E} {s : Finset ι} - (h_card_small : s.card ≤ finrank 𝕜 E + 1) - (h_inter : ∀ I ⊆ s, I.card ≤ finrank 𝕜 E + 1 → (⋂ i ∈ I, F i).Nonempty) : + (h_card_small : #s ≤ finrank 𝕜 E + 1) + (h_inter : ∀ I ⊆ s, #I ≤ finrank 𝕜 E + 1 → (⋂ i ∈ I, F i).Nonempty) : (⋂ i ∈ s, F i).Nonempty := h_inter s (by simp) h_card_small variable [FiniteDimensional 𝕜 E] @@ -78,12 +78,12 @@ If `F` is a finite family of convex sets in a vector space of finite dimension ` `k ≤ d + 1` sets of `F` intersect nontrivially, then all sets of `F` intersect nontrivially. -/ theorem helly_theorem' {F : ι → Set E} {s : Finset ι} (h_convex : ∀ i ∈ s, Convex 𝕜 (F i)) - (h_inter : ∀ I ⊆ s, I.card ≤ finrank 𝕜 E + 1 → (⋂ i ∈ I, F i).Nonempty) : + (h_inter : ∀ I ⊆ s, #I ≤ finrank 𝕜 E + 1 → (⋂ i ∈ I, F i).Nonempty) : (⋂ i ∈ s, F i).Nonempty := by classical - obtain h_card | h_card := lt_or_le s.card (finrank 𝕜 E + 1) + obtain h_card | h_card := lt_or_le #s (finrank 𝕜 E + 1) · exact helly_theorem_corner (le_of_lt h_card) h_inter - generalize hn : s.card = n + generalize hn : #s = n rw [hn] at h_card induction' n, h_card using Nat.le_induction with k h_card hk generalizing ι · exact helly_theorem_corner (le_of_eq hn) h_inter @@ -137,9 +137,9 @@ theorem helly_theorem' {F : ι → Set E} {s : Finset ι} If `F` is a family of `n` convex sets in a vector space of finite dimension `d`, with `n ≥ d + 1`, and any `d + 1` sets of `F` intersect nontrivially, then all sets of `F` intersect nontrivially. -/ theorem helly_theorem {F : ι → Set E} {s : Finset ι} - (h_card : finrank 𝕜 E + 1 ≤ s.card) + (h_card : finrank 𝕜 E + 1 ≤ #s) (h_convex : ∀ i ∈ s, Convex 𝕜 (F i)) - (h_inter : ∀ I ⊆ s, I.card = finrank 𝕜 E + 1 → (⋂ i ∈ I, F i).Nonempty) : + (h_inter : ∀ I ⊆ s, #I = finrank 𝕜 E + 1 → (⋂ i ∈ I, F i).Nonempty) : (⋂ i ∈ s, F i).Nonempty := by apply helly_theorem' h_convex intro I hI_ss hI_card @@ -153,7 +153,7 @@ If `F` is a finite set of convex sets in a vector space of finite dimension `d`, sets from `F` intersect nontrivially, then all sets from `F` intersect nontrivially. -/ theorem helly_theorem_set' {F : Finset (Set E)} (h_convex : ∀ X ∈ F, Convex 𝕜 X) - (h_inter : ∀ G : Finset (Set E), G ⊆ F → G.card ≤ finrank 𝕜 E + 1 → (⋂₀ G : Set E).Nonempty) : + (h_inter : ∀ G : Finset (Set E), G ⊆ F → #G ≤ finrank 𝕜 E + 1 → (⋂₀ G : Set E).Nonempty) : (⋂₀ (F : Set (Set E))).Nonempty := by classical -- for DecidableEq, required for the family version rw [show ⋂₀ F = ⋂ X ∈ F, (X : Set E) by ext; simp] @@ -168,9 +168,9 @@ If `F` is a finite set of convex sets in a vector space of finite dimension `d`, and any `d + 1` sets from `F` intersect nontrivially, then all sets from `F` intersect nontrivially. -/ theorem helly_theorem_set {F : Finset (Set E)} - (h_card : finrank 𝕜 E + 1 ≤ F.card) + (h_card : finrank 𝕜 E + 1 ≤ #F) (h_convex : ∀ X ∈ F, Convex 𝕜 X) - (h_inter : ∀ G : Finset (Set E), G ⊆ F → G.card = finrank 𝕜 E + 1 → (⋂₀ G : Set E).Nonempty) : + (h_inter : ∀ G : Finset (Set E), G ⊆ F → #G = finrank 𝕜 E + 1 → (⋂₀ G : Set E).Nonempty) : (⋂₀ (F : Set (Set E))).Nonempty := by apply helly_theorem_set' h_convex intro I hI_ss hI_card @@ -185,7 +185,7 @@ If `F` is a family of compact convex sets in a vector space of finite dimension `k ≤ d + 1` sets of `F` intersect nontrivially, then all sets of `F` intersect nontrivially. -/ theorem helly_theorem_compact' [TopologicalSpace E] [T2Space E] {F : ι → Set E} (h_convex : ∀ i, Convex 𝕜 (F i)) (h_compact : ∀ i, IsCompact (F i)) - (h_inter : ∀ I : Finset ι, I.card ≤ finrank 𝕜 E + 1 → (⋂ i ∈ I, F i).Nonempty) : + (h_inter : ∀ I : Finset ι, #I ≤ finrank 𝕜 E + 1 → (⋂ i ∈ I, F i).Nonempty) : (⋂ i, F i).Nonempty := by classical /- If `ι` is empty the statement is trivial. -/ @@ -214,11 +214,11 @@ then all sets of `F` intersect nontrivially. -/ theorem helly_theorem_compact [TopologicalSpace E] [T2Space E] {F : ι → Set E} (h_card : finrank 𝕜 E + 1 ≤ PartENat.card ι) (h_convex : ∀ i, Convex 𝕜 (F i)) (h_compact : ∀ i, IsCompact (F i)) - (h_inter : ∀ I : Finset ι, I.card = finrank 𝕜 E + 1 → (⋂ i ∈ I, F i).Nonempty) : + (h_inter : ∀ I : Finset ι, #I = finrank 𝕜 E + 1 → (⋂ i ∈ I, F i).Nonempty) : (⋂ i, F i).Nonempty := by apply helly_theorem_compact' h_convex h_compact intro I hI_card - have hJ : ∃ J : Finset ι, I ⊆ J ∧ J.card = finrank 𝕜 E + 1 := by + have hJ : ∃ J : Finset ι, I ⊆ J ∧ #J = finrank 𝕜 E + 1 := by by_cases h : Infinite ι · exact Infinite.exists_superset_card_eq _ _ hI_card · have : Finite ι := Finite.of_not_infinite h @@ -236,7 +236,7 @@ If `F` is a set of compact convex sets in a vector space of finite dimension `d` `k ≤ d + 1` sets from `F` intersect nontrivially, then all sets from `F` intersect nontrivially. -/ theorem helly_theorem_set_compact' [TopologicalSpace E] [T2Space E] {F : Set (Set E)} (h_convex : ∀ X ∈ F, Convex 𝕜 X) (h_compact : ∀ X ∈ F, IsCompact X) - (h_inter : ∀ G : Finset (Set E), (G : Set (Set E)) ⊆ F → G.card ≤ finrank 𝕜 E + 1 → + (h_inter : ∀ G : Finset (Set E), (G : Set (Set E)) ⊆ F → #G ≤ finrank 𝕜 E + 1 → (⋂₀ G : Set E).Nonempty) : (⋂₀ (F : Set (Set E))).Nonempty := by classical -- for DecidableEq, required for the family version @@ -259,7 +259,7 @@ then all sets from `F` intersect nontrivially. -/ theorem helly_theorem_set_compact [TopologicalSpace E] [T2Space E] {F : Set (Set E)} (h_card : finrank 𝕜 E + 1 ≤ F.encard) (h_convex : ∀ X ∈ F, Convex 𝕜 X) (h_compact : ∀ X ∈ F, IsCompact X) - (h_inter : ∀ G : Finset (Set E), (G : Set (Set E)) ⊆ F → G.card = finrank 𝕜 E + 1 → + (h_inter : ∀ G : Finset (Set E), (G : Set (Set E)) ⊆ F → #G = finrank 𝕜 E + 1 → (⋂₀ G : Set E).Nonempty) : (⋂₀ (F : Set (Set E))).Nonempty := by apply helly_theorem_set_compact' h_convex h_compact diff --git a/Mathlib/Analysis/MeanInequalities.lean b/Mathlib/Analysis/MeanInequalities.lean index ebc7bb9056c0f..d75d457d35cf7 100644 --- a/Mathlib/Analysis/MeanInequalities.lean +++ b/Mathlib/Analysis/MeanInequalities.lean @@ -494,7 +494,7 @@ theorem inner_le_Lp_mul_Lq_hasSum {f g : ι → ℝ≥0} {A B : ℝ≥0} {p q : sum of the `p`-th powers of `f i`. Version for sums over finite sets, with `ℝ≥0`-valued functions. -/ theorem rpow_sum_le_const_mul_sum_rpow (f : ι → ℝ≥0) {p : ℝ} (hp : 1 ≤ p) : - (∑ i ∈ s, f i) ^ p ≤ (card s : ℝ≥0) ^ (p - 1) * ∑ i ∈ s, f i ^ p := by + (∑ i ∈ s, f i) ^ p ≤ (#s : ℝ≥0) ^ (p - 1) * ∑ i ∈ s, f i ^ p := by cases' eq_or_lt_of_le hp with hp hp · simp [← hp] let q : ℝ := p / (p - 1) @@ -620,7 +620,7 @@ theorem inner_le_Lp_mul_Lq (hpq : IsConjExponent p q) : /-- For `1 ≤ p`, the `p`-th power of the sum of `f i` is bounded above by a constant times the sum of the `p`-th powers of `f i`. Version for sums over finite sets, with `ℝ`-valued functions. -/ theorem rpow_sum_le_const_mul_sum_rpow (hp : 1 ≤ p) : - (∑ i ∈ s, |f i|) ^ p ≤ (card s : ℝ) ^ (p - 1) * ∑ i ∈ s, |f i| ^ p := by + (∑ i ∈ s, |f i|) ^ p ≤ (#s : ℝ) ^ (p - 1) * ∑ i ∈ s, |f i| ^ p := by have := NNReal.coe_le_coe.2 (NNReal.rpow_sum_le_const_mul_sum_rpow s (fun i => ⟨_, abs_nonneg (f i)⟩) hp) @@ -722,7 +722,7 @@ theorem inner_le_Lp_mul_Lq_hasSum_of_nonneg (hpq : p.IsConjExponent q) {A B : sum of the `p`-th powers of `f i`. Version for sums over finite sets, with nonnegative `ℝ`-valued functions. -/ theorem rpow_sum_le_const_mul_sum_rpow_of_nonneg (hp : 1 ≤ p) (hf : ∀ i ∈ s, 0 ≤ f i) : - (∑ i ∈ s, f i) ^ p ≤ (card s : ℝ) ^ (p - 1) * ∑ i ∈ s, f i ^ p := by + (∑ i ∈ s, f i) ^ p ≤ (#s : ℝ) ^ (p - 1) * ∑ i ∈ s, f i ^ p := by convert rpow_sum_le_const_mul_sum_rpow s f hp using 2 <;> apply sum_congr rfl <;> intro i hi <;> simp only [abs_of_nonneg, hf i hi] diff --git a/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean b/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean index 7f88338734d62..02f606b3771ea 100644 --- a/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean +++ b/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean @@ -307,7 +307,7 @@ theorem coe_mkContinuous (C : ℝ) (H : ∀ m, ‖f m‖ ≤ C * ∏ i, ‖m i the other coordinates, then the resulting restricted function satisfies an inequality `‖f.restr v‖ ≤ C * ‖z‖^(n-k) * Π ‖v i‖` if the original function satisfies `‖f v‖ ≤ C * Π ‖v i‖`. -/ theorem restr_norm_le {k n : ℕ} (f : (MultilinearMap 𝕜 (fun _ : Fin n => G) G' : _)) - (s : Finset (Fin n)) (hk : s.card = k) (z : G) {C : ℝ} (H : ∀ m, ‖f m‖ ≤ C * ∏ i, ‖m i‖) + (s : Finset (Fin n)) (hk : #s = k) (z : G) {C : ℝ} (H : ∀ m, ‖f m‖ ≤ C * ∏ i, ‖m i‖) (v : Fin k → G) : ‖f.restr s hk z v‖ ≤ C * ‖z‖ ^ (n - k) * ∏ i, ‖v i‖ := by rw [mul_right_comm, mul_assoc] convert H _ using 2 @@ -732,12 +732,12 @@ namespace ContinuousMultilinearMap these variables, and fixing the other ones equal to a given value `z`. It is denoted by `f.restr s hk z`, where `hk` is a proof that the cardinality of `s` is `k`. The implicit identification between `Fin k` and `s` that we use is the canonical (increasing) bijection. -/ -def restr {k n : ℕ} (f : (G[×n]→L[𝕜] G' : _)) (s : Finset (Fin n)) (hk : s.card = k) (z : G) : +def restr {k n : ℕ} (f : (G[×n]→L[𝕜] G' : _)) (s : Finset (Fin n)) (hk : #s = k) (z : G) : G[×k]→L[𝕜] G' := (f.toMultilinearMap.restr s hk z).mkContinuous (‖f‖ * ‖z‖ ^ (n - k)) fun _ => MultilinearMap.restr_norm_le _ _ _ _ f.le_opNorm _ -theorem norm_restr {k n : ℕ} (f : G[×n]→L[𝕜] G') (s : Finset (Fin n)) (hk : s.card = k) (z : G) : +theorem norm_restr {k n : ℕ} (f : G[×n]→L[𝕜] G') (s : Finset (Fin n)) (hk : #s = k) (z : G) : ‖f.restr s hk z‖ ≤ ‖f‖ * ‖z‖ ^ (n - k) := by apply MultilinearMap.mkContinuous_norm_le exact mul_nonneg (norm_nonneg _) (pow_nonneg (norm_nonneg _) _) diff --git a/Mathlib/Analysis/NormedSpace/Multilinear/Curry.lean b/Mathlib/Analysis/NormedSpace/Multilinear/Curry.lean index 2ce89faee16da..3e81916160893 100644 --- a/Mathlib/Analysis/NormedSpace/Multilinear/Curry.lean +++ b/Mathlib/Analysis/NormedSpace/Multilinear/Curry.lean @@ -560,7 +560,7 @@ variable (𝕜 G G') {k l : ℕ} {s : Finset (Fin n)} `l`, then the space of continuous multilinear maps `G [×n]→L[𝕜] G'` of `n` variables is isomorphic to the space of continuous multilinear maps `G [×k]→L[𝕜] G [×l]→L[𝕜] G'` of `k` variables taking values in the space of continuous multilinear maps of `l` variables. -/ -def curryFinFinset {k l n : ℕ} {s : Finset (Fin n)} (hk : s.card = k) (hl : sᶜ.card = l) : +def curryFinFinset {k l n : ℕ} {s : Finset (Fin n)} (hk : #s = k) (hl : #sᶜ = l) : (G[×n]→L[𝕜] G') ≃ₗᵢ[𝕜] G[×k]→L[𝕜] G[×l]→L[𝕜] G' := (domDomCongrₗᵢ 𝕜 G G' (finSumEquivOfFinset hk hl).symm).trans (currySumEquiv 𝕜 (Fin k) (Fin l) G G') @@ -568,31 +568,31 @@ def curryFinFinset {k l n : ℕ} {s : Finset (Fin n)} (hk : s.card = k) (hl : s variable {𝕜 G G'} @[simp] -theorem curryFinFinset_apply (hk : s.card = k) (hl : sᶜ.card = l) (f : G[×n]→L[𝕜] G') +theorem curryFinFinset_apply (hk : #s = k) (hl : #sᶜ = l) (f : G[×n]→L[𝕜] G') (mk : Fin k → G) (ml : Fin l → G) : curryFinFinset 𝕜 G G' hk hl f mk ml = f fun i => Sum.elim mk ml ((finSumEquivOfFinset hk hl).symm i) := rfl @[simp] -theorem curryFinFinset_symm_apply (hk : s.card = k) (hl : sᶜ.card = l) +theorem curryFinFinset_symm_apply (hk : #s = k) (hl : #sᶜ = l) (f : G[×k]→L[𝕜] G[×l]→L[𝕜] G') (m : Fin n → G) : (curryFinFinset 𝕜 G G' hk hl).symm f m = f (fun i => m <| finSumEquivOfFinset hk hl (Sum.inl i)) fun i => m <| finSumEquivOfFinset hk hl (Sum.inr i) := rfl -theorem curryFinFinset_symm_apply_piecewise_const (hk : s.card = k) (hl : sᶜ.card = l) +theorem curryFinFinset_symm_apply_piecewise_const (hk : #s = k) (hl : #sᶜ = l) (f : G[×k]→L[𝕜] G[×l]→L[𝕜] G') (x y : G) : (curryFinFinset 𝕜 G G' hk hl).symm f (s.piecewise (fun _ => x) fun _ => y) = f (fun _ => x) fun _ => y := MultilinearMap.curryFinFinset_symm_apply_piecewise_const hk hl _ x y @[simp] -theorem curryFinFinset_symm_apply_const (hk : s.card = k) (hl : sᶜ.card = l) +theorem curryFinFinset_symm_apply_const (hk : #s = k) (hl : #sᶜ = l) (f : G[×k]→L[𝕜] G[×l]→L[𝕜] G') (x : G) : ((curryFinFinset 𝕜 G G' hk hl).symm f fun _ => x) = f (fun _ => x) fun _ => x := rfl -theorem curryFinFinset_apply_const (hk : s.card = k) (hl : sᶜ.card = l) (f : G[×n]→L[𝕜] G') +theorem curryFinFinset_apply_const (hk : #s = k) (hl : #sᶜ = l) (f : G[×n]→L[𝕜] G') (x y : G) : (curryFinFinset 𝕜 G G' hk hl f (fun _ => x) fun _ => y) = f (s.piecewise (fun _ => x) fun _ => y) := by refine (curryFinFinset_symm_apply_piecewise_const hk hl _ _ _).symm.trans ?_ diff --git a/Mathlib/Analysis/SpecificLimits/FloorPow.lean b/Mathlib/Analysis/SpecificLimits/FloorPow.lean index 0d410a3b11645..6d020f15e36d5 100644 --- a/Mathlib/Analysis/SpecificLimits/FloorPow.lean +++ b/Mathlib/Analysis/SpecificLimits/FloorPow.lean @@ -216,7 +216,7 @@ theorem tendsto_div_of_monotone_of_tendsto_div_floor_pow (u : ℕ → ℝ) (l : /-- The sum of `1/(c^i)^2` above a threshold `j` is comparable to `1/j^2`, up to a multiplicative constant. -/ theorem sum_div_pow_sq_le_div_sq (N : ℕ) {j : ℝ} (hj : 0 < j) {c : ℝ} (hc : 1 < c) : - (∑ i ∈ (range N).filter (j < c ^ ·), (1 : ℝ) / (c ^ i) ^ 2) ≤ c ^ 3 * (c - 1)⁻¹ / j ^ 2 := by + (∑ i ∈ range N with j < c ^ i, (1 : ℝ) / (c ^ i) ^ 2) ≤ c ^ 3 * (c - 1)⁻¹ / j ^ 2 := by have cpos : 0 < c := zero_lt_one.trans hc have A : (0 : ℝ) < c⁻¹ ^ 2 := sq_pos_of_pos (inv_pos.2 cpos) have B : c ^ 2 * ((1 : ℝ) - c⁻¹ ^ 2)⁻¹ ≤ c ^ 3 * (c - 1)⁻¹ := by @@ -229,7 +229,7 @@ theorem sum_div_pow_sq_le_div_sq (N : ℕ) {j : ℝ} (hj : 0 < j) {c : ℝ} (hc 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_of_one_lt₀ hc) two_ne_zero calc - (∑ i ∈ (range N).filter (j < c ^ ·), (1 : ℝ) / (c ^ i) ^ 2) ≤ + (∑ i ∈ range N with j < c ^ i, (1 : ℝ) / (c ^ i) ^ 2) ≤ ∑ i ∈ Ico ⌊Real.log j / Real.log c⌋₊ N, (1 : ℝ) / (c ^ i) ^ 2 := by refine sum_le_sum_of_subset_of_nonneg (fun i hi ↦ ?_) (by intros; positivity) simp only [mem_filter, mem_range] at hi @@ -276,17 +276,16 @@ theorem mul_pow_le_nat_floor_pow {c : ℝ} (hc : 1 < c) (i : ℕ) : (1 - c⁻¹) /-- The sum of `1/⌊c^i⌋₊^2` above a threshold `j` is comparable to `1/j^2`, up to a multiplicative constant. -/ theorem sum_div_nat_floor_pow_sq_le_div_sq (N : ℕ) {j : ℝ} (hj : 0 < j) {c : ℝ} (hc : 1 < c) : - (∑ i ∈ (range N).filter (j < ⌊c ^ ·⌋₊), (1 : ℝ) / (⌊c ^ i⌋₊ : ℝ) ^ 2) ≤ + (∑ i ∈ range N with j < ⌊c ^ i⌋₊, (1 : ℝ) / (⌊c ^ i⌋₊ : ℝ) ^ 2) ≤ c ^ 5 * (c - 1)⁻¹ ^ 3 / j ^ 2 := by have cpos : 0 < c := zero_lt_one.trans hc have A : 0 < 1 - c⁻¹ := sub_pos.2 (inv_lt_one_of_one_lt₀ hc) calc - (∑ i ∈ (range N).filter (j < ⌊c ^ ·⌋₊), (1 : ℝ) / (⌊c ^ i⌋₊ : ℝ) ^ 2) ≤ - ∑ i ∈ (range N).filter (j < c ^ ·), (1 : ℝ) / (⌊c ^ i⌋₊ : ℝ) ^ 2 := by - apply sum_le_sum_of_subset_of_nonneg - · exact monotone_filter_right _ fun k hk ↦ hk.trans_le <| Nat.floor_le (by positivity) - · intros; positivity - _ ≤ ∑ i ∈ (range N).filter (j < c ^ ·), (1 - c⁻¹)⁻¹ ^ 2 * ((1 : ℝ) / (c ^ i) ^ 2) := by + (∑ i ∈ range N with j < ⌊c ^ i⌋₊, (1 : ℝ) / (⌊c ^ i⌋₊ : ℝ) ^ 2) ≤ + ∑ i ∈ range N with j < c ^ i, (1 : ℝ) / (⌊c ^ i⌋₊ : ℝ) ^ 2 := by + gcongr + exact fun k hk ↦ hk.trans_le <| Nat.floor_le (by positivity) + _ ≤ ∑ i ∈ range N with j < c ^ i, (1 - c⁻¹)⁻¹ ^ 2 * ((1 : ℝ) / (c ^ i) ^ 2) := by refine sum_le_sum fun i _hi => ?_ rw [mul_div_assoc', mul_one, div_le_div_iff]; rotate_left · apply sq_pos_of_pos diff --git a/Mathlib/Computability/AkraBazzi/AkraBazzi.lean b/Mathlib/Computability/AkraBazzi/AkraBazzi.lean index 5f7d1c28e9623..9cfbcd713b62a 100644 --- a/Mathlib/Computability/AkraBazzi/AkraBazzi.lean +++ b/Mathlib/Computability/AkraBazzi/AkraBazzi.lean @@ -644,9 +644,9 @@ lemma eventually_atTop_sumTransform_le : _ ≤ u := by exact_mod_cast hu'.1 _ ≤ n ^ (p a b) * (∑ _u ∈ Finset.Ico (r i n) n, c₂ * g n / (r i n) ^ ((p a b) + 1)) := by gcongr with u hu; rw [Finset.mem_Ico] at hu; exact hu.1 - _ ≤ n ^ (p a b) * (Finset.Ico (r i n) n).card • (c₂ * g n / (r i n) ^ ((p a b) + 1)) := by + _ ≤ n ^ p a b * #(Ico (r i n) n) • (c₂ * g n / r i n ^ (p a b + 1)) := by gcongr; exact Finset.sum_le_card_nsmul _ _ _ (fun x _ => by rfl) - _ = n ^ (p a b) * (Finset.Ico (r i n) n).card * (c₂ * g n / (r i n) ^ ((p a b) + 1)) := by + _ = n ^ p a b * #(Ico (r i n) n) * (c₂ * g n / r i n ^ (p a b + 1)) := by rw [nsmul_eq_mul, mul_assoc] _ = n ^ (p a b) * (n - r i n) * (c₂ * g n / (r i n) ^ ((p a b) + 1)) := by congr; rw [Nat.card_Ico, Nat.cast_sub (le_of_lt <| hr_lt_n i)] @@ -682,9 +682,9 @@ lemma eventually_atTop_sumTransform_le : _ ≤ u := by exact hu.1 exact rpow_le_rpow_of_exponent_nonpos (by positivity) (by exact_mod_cast (le_of_lt hu.2)) (le_of_lt hp) - _ ≤ n ^ (p a b) * (Finset.Ico (r i n) n).card • (c₂ * g n / n ^ ((p a b) + 1)) := by + _ ≤ n ^ p a b * #(Ico (r i n) n) • (c₂ * g n / n ^ (p a b + 1)) := by gcongr; exact Finset.sum_le_card_nsmul _ _ _ (fun x _ => by rfl) - _ = n ^ (p a b) * (Finset.Ico (r i n) n).card * (c₂ * g n / n ^ ((p a b) + 1)) := by + _ = n ^ p a b * #(Ico (r i n) n) * (c₂ * g n / n ^ (p a b + 1)) := by rw [nsmul_eq_mul, mul_assoc] _ = n ^ (p a b) * (n - r i n) * (c₂ * g n / n ^ ((p a b) + 1)) := by congr; rw [Nat.card_Ico, Nat.cast_sub (le_of_lt <| hr_lt_n i)] @@ -729,9 +729,9 @@ lemma eventually_atTop_sumTransform_ge : positivity · rw [Finset.mem_Ico] at hu exact le_of_lt hu.2 - _ ≥ n ^ (p a b) * (Finset.Ico (r i n) n).card • (c₂ * g n / n ^ ((p a b) + 1)) := by + _ ≥ n ^ p a b * #(Ico (r i n) n) • (c₂ * g n / n ^ (p a b + 1)) := by gcongr; exact Finset.card_nsmul_le_sum _ _ _ (fun x _ => by rfl) - _ = n ^ (p a b) * (Finset.Ico (r i n) n).card * (c₂ * g n / n ^ ((p a b) + 1)) := by + _ = n ^ p a b * #(Ico (r i n) n) * (c₂ * g n / n ^ (p a b + 1)) := by rw [nsmul_eq_mul, mul_assoc] _ = n ^ (p a b) * (n - r i n) * (c₂ * g n / n ^ ((p a b) + 1)) := by congr; rw [Nat.card_Ico, Nat.cast_sub (le_of_lt <| hr_lt_n i)] @@ -764,12 +764,12 @@ lemma eventually_atTop_sumTransform_ge : · rw [Finset.mem_Ico] at hu exact rpow_le_rpow_of_exponent_nonpos (by positivity) (by exact_mod_cast hu.1) (le_of_lt hp) - _ ≥ n ^ (p a b) * (Finset.Ico (r i n) n).card • (c₂ * g n / (r i n) ^ ((p a b) + 1)) := by + _ ≥ n ^ p a b * #(Ico (r i n) n) • (c₂ * g n / r i n ^ (p a b + 1)) := by gcongr; exact Finset.card_nsmul_le_sum _ _ _ (fun x _ => by rfl) - _ = n ^ (p a b) * (Finset.Ico (r i n) n).card * (c₂ * g n / (r i n) ^ ((p a b) + 1)) := by + _ = n ^ p a b * #(Ico (r i n) n) * (c₂ * g n / r i n ^ (p a b + 1)) := by rw [nsmul_eq_mul, mul_assoc] - _ ≥ n ^ (p a b) * (Finset.Ico (r i n) n).card * (c₂ * g n / (c₁ * n) ^ ((p a b) + 1)) := by - gcongr n^(p a b) * (Finset.Ico (r i n) n).card * (c₂ * g n / ?_) + _ ≥ n ^ p a b * #(Ico (r i n) n) * (c₂ * g n / (c₁ * n) ^ (p a b + 1)) := by + gcongr n ^ p a b * #(Ico (r i n) n) * (c₂ * g n / ?_) exact rpow_le_rpow_of_exponent_nonpos (by positivity) (hn₁ i) (le_of_lt hp) _ = n ^ (p a b) * (n - r i n) * (c₂ * g n / (c₁ * n) ^ ((p a b) + 1)) := by congr; rw [Nat.card_Ico, Nat.cast_sub (le_of_lt <| hr_lt_n i)] diff --git a/Mathlib/FieldTheory/Finite/Basic.lean b/Mathlib/FieldTheory/Finite/Basic.lean index 715387ba0669b..f688f79dda165 100644 --- a/Mathlib/FieldTheory/Finite/Basic.lean +++ b/Mathlib/FieldTheory/Finite/Basic.lean @@ -61,10 +61,10 @@ open Polynomial /-- The cardinality of a field is at most `n` times the cardinality of the image of a degree `n` polynomial -/ theorem card_image_polynomial_eval [DecidableEq R] [Fintype R] {p : R[X]} (hp : 0 < p.degree) : - Fintype.card R ≤ natDegree p * (univ.image fun x => eval x p).card := + Fintype.card R ≤ natDegree p * #(univ.image fun x => eval x p) := Finset.card_le_mul_card_image _ _ (fun a _ => calc - _ = (p - C a).roots.toFinset.card := + _ = #(p - C a).roots.toFinset := congr_arg card (by simp [Finset.ext_iff, ← mem_roots_sub_C hp]) _ ≤ Multiset.card (p - C a).roots := Multiset.toFinset_card_le _ _ ≤ _ := card_roots_sub_C' hp) @@ -80,17 +80,17 @@ theorem exists_root_sum_quadratic [Fintype R] {f g : R[X]} (hf2 : degree f = 2) rcases this with ⟨x, ⟨a, _, ha⟩, ⟨b, _, hb⟩⟩ exact ⟨a, b, by rw [ha, ← hb, eval_neg, neg_add_cancel]⟩ fun hd : Disjoint _ _ => - lt_irrefl (2 * ((univ.image fun x : R => eval x f) ∪ univ.image fun x : R => eval x (-g)).card) <| - calc 2 * ((univ.image fun x : R => eval x f) ∪ univ.image fun x : R => eval x (-g)).card + lt_irrefl (2 * #((univ.image fun x : R => eval x f) ∪ univ.image fun x : R => eval x (-g))) <| + calc 2 * #((univ.image fun x : R => eval x f) ∪ univ.image fun x : R => eval x (-g)) ≤ 2 * Fintype.card R := Nat.mul_le_mul_left _ (Finset.card_le_univ _) _ = Fintype.card R + Fintype.card R := two_mul _ - _ < natDegree f * (univ.image fun x : R => eval x f).card + - natDegree (-g) * (univ.image fun x : R => eval x (-g)).card := + _ < natDegree f * #(univ.image fun x : R => eval x f) + + natDegree (-g) * #(univ.image fun x : R => eval x (-g)) := (add_lt_add_of_lt_of_le (lt_of_le_of_ne (card_image_polynomial_eval (by rw [hf2]; decide)) (mt (congr_arg (· % 2)) (by simp [natDegree_eq_of_degree_eq_some hf2, hR]))) (card_image_polynomial_eval (by rw [degree_neg, hg2]; decide))) - _ = 2 * ((univ.image fun x : R => eval x f) ∪ univ.image fun x : R => eval x (-g)).card := by + _ = 2 * #((univ.image fun x : R => eval x f) ∪ univ.image fun x : R => eval x (-g)) := by rw [card_union_of_disjoint hd] simp [natDegree_eq_of_degree_eq_some hf2, natDegree_eq_of_degree_eq_some hg2, mul_add] diff --git a/Mathlib/Geometry/Euclidean/Circumcenter.lean b/Mathlib/Geometry/Euclidean/Circumcenter.lean index ab715d6d83899..bd9d26e0df104 100644 --- a/Mathlib/Geometry/Euclidean/Circumcenter.lean +++ b/Mathlib/Geometry/Euclidean/Circumcenter.lean @@ -453,7 +453,7 @@ theorem orthogonalProjection_eq_circumcenter_of_dist_eq {n : ℕ} (s : Simplex /-- The orthogonal projection of the circumcenter onto a face is the circumcenter of that face. -/ theorem orthogonalProjection_circumcenter {n : ℕ} (s : Simplex ℝ P n) {fs : Finset (Fin (n + 1))} - {m : ℕ} (h : fs.card = m + 1) : + {m : ℕ} (h : #fs = m + 1) : ↑((s.face h).orthogonalProjectionSpan s.circumcenter) = (s.face h).circumcenter := haveI hr : ∃ r, ∀ i, dist ((s.face h).points i) s.circumcenter = r := by use s.circumradius @@ -564,7 +564,7 @@ theorem point_eq_affineCombination_of_pointsWithCircumcenter {n : ℕ} (s : Simp terms of `pointsWithCircumcenter`. -/ def centroidWeightsWithCircumcenter {n : ℕ} (fs : Finset (Fin (n + 1))) : PointsWithCircumcenterIndex n → ℝ - | pointIndex i => if i ∈ fs then (card fs : ℝ)⁻¹ else 0 + | pointIndex i => if i ∈ fs then (#fs : ℝ)⁻¹ else 0 | circumcenterIndex => 0 /-- `centroidWeightsWithCircumcenter` sums to 1, if the `Finset` is nonempty. -/ @@ -584,7 +584,7 @@ theorem centroid_eq_affineCombination_of_pointsWithCircumcenter {n : ℕ} (s : S simp_rw [centroid_def, affineCombination_apply, weightedVSubOfPoint_apply, sum_pointsWithCircumcenter, centroidWeightsWithCircumcenter, pointsWithCircumcenter_point, zero_smul, add_zero, centroidWeights, - ← sum_indicator_subset_of_eq_zero (Function.const (Fin (n + 1)) (card fs : ℝ)⁻¹) + ← sum_indicator_subset_of_eq_zero (Function.const (Fin (n + 1)) (#fs : ℝ)⁻¹) (fun i wi => wi • (s.points i -ᵥ Classical.choice AddTorsor.nonempty)) fs.subset_univ fun _ => zero_smul ℝ _, Set.indicator_apply] @@ -638,7 +638,7 @@ theorem reflection_circumcenter_eq_affineCombination_of_pointsWithCircumcenter { reflection (affineSpan ℝ (s.points '' {i₁, i₂})) s.circumcenter = (univ : Finset (PointsWithCircumcenterIndex n)).affineCombination ℝ s.pointsWithCircumcenter (reflectionCircumcenterWeightsWithCircumcenter i₁ i₂) := by - have hc : card ({i₁, i₂} : Finset (Fin (n + 1))) = 2 := by simp [h] + have hc : #{i₁, i₂} = 2 := by simp [h] -- Making the next line a separate definition helps the elaborator: set W : AffineSubspace ℝ P := affineSpan ℝ (s.points '' {i₁, i₂}) have h_faces : diff --git a/Mathlib/Geometry/Euclidean/MongePoint.lean b/Mathlib/Geometry/Euclidean/MongePoint.lean index d829d76dfb4ef..f08ceefb2ff48 100644 --- a/Mathlib/Geometry/Euclidean/MongePoint.lean +++ b/Mathlib/Geometry/Euclidean/MongePoint.lean @@ -168,7 +168,7 @@ theorem mongePointVSubFaceCentroidWeightsWithCircumcenter_eq_sub {n : ℕ} {i₁ cases' i with i · rw [Pi.sub_apply, mongePointWeightsWithCircumcenter, centroidWeightsWithCircumcenter, mongePointVSubFaceCentroidWeightsWithCircumcenter] - have hu : card ({i₁, i₂}ᶜ : Finset (Fin (n + 3))) = n + 1 := by + have hu : #{i₁, i₂}ᶜ = n + 1 := by simp [card_compl, Fintype.card_fin, h] rw [hu] by_cases hi : i = i₁ ∨ i = i₂ <;> simp [compl_eq_univ_sdiff, hi] @@ -306,7 +306,7 @@ theorem eq_mongePoint_of_forall_mem_mongePlane {n : ℕ} {s : Simplex ℝ P (n + rw [hu, ← vectorSpan_image_eq_span_vsub_set_left_ne ℝ _ (Set.mem_univ _), Set.image_univ] at hi have hv : p -ᵥ s.mongePoint ∈ vectorSpan ℝ (Set.range s.points) := by let s₁ : Finset (Fin (n + 3)) := univ.erase i₁ - obtain ⟨i₂, h₂⟩ := card_pos.1 (show 0 < card s₁ by simp [s₁, card_erase_of_mem]) + obtain ⟨i₂, h₂⟩ := card_pos.1 (show 0 < #s₁ by simp [s₁, card_erase_of_mem]) have h₁₂ : i₁ ≠ i₂ := (ne_of_mem_erase h₂).symm exact (Submodule.mem_inf.1 (h' i₂ h₁₂)).2 exact Submodule.disjoint_def.1 (vectorSpan ℝ (Set.range s.points)).orthogonal_disjoint _ hv hi @@ -359,7 +359,7 @@ theorem finrank_direction_altitude {n : ℕ} (s : Simplex ℝ P (n + 1)) (i : Fi rw [direction_altitude] have h := Submodule.finrank_add_inf_finrank_orthogonal (vectorSpan_mono ℝ (Set.image_subset_range s.points ↑(univ.erase i))) - have hc : card (univ.erase i) = n + 1 := by rw [card_erase_of_mem (mem_univ _)]; simp + have hc : #(univ.erase i) = n + 1 := by rw [card_erase_of_mem (mem_univ _)]; simp refine add_left_cancel (_root_.trans h ?_) classical rw [s.independent.finrank_vectorSpan (Fintype.card_fin _), ← Finset.coe_image, diff --git a/Mathlib/GroupTheory/Index.lean b/Mathlib/GroupTheory/Index.lean index 237170a95e0dc..66ce7103de92b 100644 --- a/Mathlib/GroupTheory/Index.lean +++ b/Mathlib/GroupTheory/Index.lean @@ -598,7 +598,7 @@ variable {G M F : Type*} [Group G] [Fintype G] [Monoid M] [DecidableEq M] @[to_additive] lemma card_fiber_eq_of_mem_range (f : F) {x y : M} (hx : x ∈ Set.range f) (hy : y ∈ Set.range f) : - (univ.filter <| fun g => f g = x).card = (univ.filter <| fun g => f g = y).card := by + #{g | f g = x} = #{g | f g = y} := by rcases hx with ⟨x, rfl⟩ rcases hy with ⟨y, rfl⟩ rcases mul_left_surjective x y with ⟨y, rfl⟩ diff --git a/Mathlib/GroupTheory/Perm/Cycle/Basic.lean b/Mathlib/GroupTheory/Perm/Cycle/Basic.lean index 0fc24ba0a1006..6cf77b57dbfc2 100644 --- a/Mathlib/GroupTheory/Perm/Cycle/Basic.lean +++ b/Mathlib/GroupTheory/Perm/Cycle/Basic.lean @@ -310,7 +310,7 @@ protected theorem IsSwap.isCycle : IsSwap f → IsCycle f := by variable [Fintype α] -theorem IsCycle.two_le_card_support (h : IsCycle f) : 2 ≤ f.support.card := +theorem IsCycle.two_le_card_support (h : IsCycle f) : 2 ≤ #f.support := two_le_card_support_of_ne_one h.ne_one /-- The subgroup generated by a cycle is in bijection with its support -/ @@ -351,7 +351,7 @@ theorem IsCycle.zpowersEquivSupport_symm_apply {σ : Perm α} (hσ : IsCycle σ) ⟨σ ^ n, n, rfl⟩ := (Equiv.symm_apply_eq _).2 hσ.zpowersEquivSupport_apply -protected theorem IsCycle.orderOf (hf : IsCycle f) : orderOf f = f.support.card := by +protected theorem IsCycle.orderOf (hf : IsCycle f) : orderOf f = #f.support := by rw [← Fintype.card_zpowers, ← Fintype.card_coe] convert Fintype.card_congr (IsCycle.zpowersEquivSupport hf) @@ -438,12 +438,12 @@ theorem IsCycle.swap_mul {α : Type*} [DecidableEq α] {f : Perm α} (hf : IsCyc isCycle_swap_mul_aux₂ (i - 1) hy hi⟩ -theorem IsCycle.sign {f : Perm α} (hf : IsCycle f) : sign f = -(-1) ^ f.support.card := +theorem IsCycle.sign {f : Perm α} (hf : IsCycle f) : sign f = -(-1) ^ #f.support := let ⟨x, hx⟩ := hf calc Perm.sign f = Perm.sign (swap x (f x) * (swap x (f x) * f)) := by {rw [← mul_assoc, mul_def, mul_def, swap_swap, trans_refl]} - _ = -(-1) ^ f.support.card := + _ = -(-1) ^ #f.support := if h1 : f (f x) = x then by have h : swap x (f x) * f = 1 := by simp only [mul_def, one_def] @@ -452,14 +452,13 @@ theorem IsCycle.sign {f : Perm α} (hf : IsCycle f) : sign f = -(-1) ^ f.support hf.eq_swap_of_apply_apply_eq_self hx.1 h1, card_support_swap hx.1.symm] rfl else by - have h : card (support (swap x (f x) * f)) + 1 = card (support f) := by + have h : #(swap x (f x) * f).support + 1 = #f.support := by rw [← insert_erase (mem_support.2 hx.1), support_swap_mul_eq _ _ h1, card_insert_of_not_mem (not_mem_erase _ _), sdiff_singleton_eq_erase] - have : card (support (swap x (f x) * f)) < card (support f) := - card_support_swap_mul hx.1 + have : #(swap x (f x) * f).support < #f.support := card_support_swap_mul hx.1 rw [sign_mul, sign_swap hx.1.symm, (hf.swap_mul hx.1 h1).sign, ← h] simp only [mul_neg, neg_mul, one_mul, neg_neg, pow_add, pow_one, mul_one] -termination_by f.support.card +termination_by #f.support theorem IsCycle.of_pow {n : ℕ} (h1 : IsCycle (f ^ n)) (h2 : f.support ⊆ (f ^ n).support) : IsCycle f := by @@ -472,7 +471,7 @@ theorem IsCycle.of_pow {n : ℕ} (h1 : IsCycle (f ^ n)) (h2 : f.support ⊆ (f ^ exact ⟨n * i, by rwa [zpow_mul]⟩ -- The lemma `support_zpow_le` is relevant. It means that `h2` is equivalent to --- `σ.support = (σ ^ n).support`, as well as to `σ.support.card ≤ (σ ^ n).support.card`. +-- `σ.support = (σ ^ n).support`, as well as to `#σ.support ≤ #(σ ^ n).support`. theorem IsCycle.of_zpow {n : ℤ} (h1 : IsCycle (f ^ n)) (h2 : f.support ⊆ (f ^ n).support) : IsCycle f := by cases n @@ -645,7 +644,7 @@ section Conjugation variable [Fintype α] [DecidableEq α] {σ τ : Perm α} -theorem IsCycle.isConj (hσ : IsCycle σ) (hτ : IsCycle τ) (h : σ.support.card = τ.support.card) : +theorem IsCycle.isConj (hσ : IsCycle σ) (hτ : IsCycle τ) (h : #σ.support = #τ.support) : IsConj σ τ := by refine isConj_of_support_equiv @@ -665,7 +664,7 @@ theorem IsCycle.isConj (hσ : IsCycle σ) (hτ : IsCycle τ) (h : σ.support.car rfl theorem IsCycle.isConj_iff (hσ : IsCycle σ) (hτ : IsCycle τ) : - IsConj σ τ ↔ σ.support.card = τ.support.card where + IsConj σ τ ↔ #σ.support = #τ.support where mp h := by obtain ⟨π, rfl⟩ := (_root_.isConj_iff).1 h refine Finset.card_bij (fun a _ => π a) (fun _ ha => ?_) (fun _ _ _ _ ab => π.injective ab) @@ -770,16 +769,15 @@ protected theorem IsCycleOn.subtypePerm (hf : f.IsCycleOn s) : -- TODO: Theory of order of an element under an action theorem IsCycleOn.pow_apply_eq {s : Finset α} (hf : f.IsCycleOn s) (ha : a ∈ s) {n : ℕ} : - (f ^ n) a = a ↔ s.card ∣ n := by + (f ^ n) a = a ↔ #s ∣ n := by obtain rfl | hs := Finset.eq_singleton_or_nontrivial ha · rw [coe_singleton, isCycleOn_singleton] at hf simpa using IsFixedPt.iterate hf n classical - have h : ∀ x ∈ s.attach, ¬f ↑x = ↑x := fun x _ => hf.apply_ne hs x.2 + have h (x : s) : ¬f x = x := hf.apply_ne hs x.2 have := (hf.isCycle_subtypePerm hs).orderOf - simp only [coe_sort_coe, support_subtype_perm, ne_eq, decide_not, Bool.not_eq_true', - decide_eq_false_iff_not, mem_attach, forall_true_left, Subtype.forall, filter_true_of_mem h, - card_attach] at this + simp only [coe_sort_coe, support_subtype_perm, ne_eq, h, not_false_eq_true, univ_eq_attach, + mem_attach, imp_self, implies_true, filter_true_of_mem, card_attach] at this rw [← this, orderOf_dvd_iff_pow_eq_one, (hf.isCycle_subtypePerm hs).pow_eq_one_iff' (ne_of_apply_ne ((↑) : s → α) <| hf.apply_ne hs (⟨a, ha⟩ : s).2)] @@ -789,32 +787,32 @@ theorem IsCycleOn.pow_apply_eq {s : Finset α} (hf : f.IsCycleOn s) (ha : a ∈ simp theorem IsCycleOn.zpow_apply_eq {s : Finset α} (hf : f.IsCycleOn s) (ha : a ∈ s) : - ∀ {n : ℤ}, (f ^ n) a = a ↔ (s.card : ℤ) ∣ n + ∀ {n : ℤ}, (f ^ n) a = a ↔ (#s : ℤ) ∣ n | Int.ofNat _ => (hf.pow_apply_eq ha).trans Int.natCast_dvd_natCast.symm | Int.negSucc n => by rw [zpow_negSucc, ← inv_pow] exact (hf.inv.pow_apply_eq ha).trans (dvd_neg.trans Int.natCast_dvd_natCast).symm theorem IsCycleOn.pow_apply_eq_pow_apply {s : Finset α} (hf : f.IsCycleOn s) (ha : a ∈ s) - {m n : ℕ} : (f ^ m) a = (f ^ n) a ↔ m ≡ n [MOD s.card] := by + {m n : ℕ} : (f ^ m) a = (f ^ n) a ↔ m ≡ n [MOD #s] := by rw [Nat.modEq_iff_dvd, ← hf.zpow_apply_eq ha] simp [sub_eq_neg_add, zpow_add, eq_inv_iff_eq, eq_comm] theorem IsCycleOn.zpow_apply_eq_zpow_apply {s : Finset α} (hf : f.IsCycleOn s) (ha : a ∈ s) - {m n : ℤ} : (f ^ m) a = (f ^ n) a ↔ m ≡ n [ZMOD s.card] := by + {m n : ℤ} : (f ^ m) a = (f ^ n) a ↔ m ≡ n [ZMOD #s] := by rw [Int.modEq_iff_dvd, ← hf.zpow_apply_eq ha] simp [sub_eq_neg_add, zpow_add, eq_inv_iff_eq, eq_comm] theorem IsCycleOn.pow_card_apply {s : Finset α} (hf : f.IsCycleOn s) (ha : a ∈ s) : - (f ^ s.card) a = a := + (f ^ #s) a = a := (hf.pow_apply_eq ha).2 dvd_rfl theorem IsCycleOn.exists_pow_eq {s : Finset α} (hf : f.IsCycleOn s) (ha : a ∈ s) (hb : b ∈ s) : - ∃ n < s.card, (f ^ n) a = b := by + ∃ n < #s, (f ^ n) a = b := by classical obtain ⟨n, rfl⟩ := hf.2 ha hb - obtain ⟨k, hk⟩ := (Int.mod_modEq n s.card).symm.dvd - refine ⟨n.natMod s.card, Int.natMod_lt (Nonempty.card_pos ⟨a, ha⟩).ne', ?_⟩ + obtain ⟨k, hk⟩ := (Int.mod_modEq n #s).symm.dvd + refine ⟨n.natMod #s, Int.natMod_lt (Nonempty.card_pos ⟨a, ha⟩).ne', ?_⟩ rw [← zpow_natCast, Int.natMod, Int.toNat_of_nonneg (Int.emod_nonneg _ <| Nat.cast_ne_zero.2 (Nonempty.card_pos ⟨a, ha⟩).ne'), sub_eq_iff_eq_add'.1 hk, zpow_add, zpow_mul] @@ -931,7 +929,7 @@ namespace Finset variable {f : Perm α} {s : Finset α} theorem product_self_eq_disjiUnion_perm_aux (hf : f.IsCycleOn s) : - (range s.card : Set ℕ).PairwiseDisjoint fun k => + (range #s : Set ℕ).PairwiseDisjoint fun k => s.map ⟨fun i => (i, (f ^ k) i), fun _ _ => congr_arg Prod.fst⟩ := by obtain hs | _ := (s : Set α).subsingleton_or_nontrivial · refine Set.Subsingleton.pairwise ?_ _ @@ -959,7 +957,7 @@ The diagonals are given by the cycle `f`. -/ theorem product_self_eq_disjiUnion_perm (hf : f.IsCycleOn s) : s ×ˢ s = - (range s.card).disjiUnion + (range #s).disjiUnion (fun k => s.map ⟨fun i => (i, (f ^ k) i), fun _ _ => congr_arg Prod.fst⟩) (product_self_eq_disjiUnion_perm_aux hf) := by ext ⟨a, b⟩ @@ -978,12 +976,12 @@ namespace Finset variable [Semiring α] [AddCommMonoid β] [Module α β] {s : Finset ι} {σ : Perm ι} theorem sum_smul_sum_eq_sum_perm (hσ : σ.IsCycleOn s) (f : ι → α) (g : ι → β) : - (∑ i ∈ s, f i) • ∑ i ∈ s, g i = ∑ k ∈ range s.card, ∑ i ∈ s, f i • g ((σ ^ k) i) := by + (∑ i ∈ s, f i) • ∑ i ∈ s, g i = ∑ k ∈ range #s, ∑ i ∈ s, f i • g ((σ ^ k) i) := by rw [sum_smul_sum, ← sum_product'] simp_rw [product_self_eq_disjiUnion_perm hσ, sum_disjiUnion, sum_map, Embedding.coeFn_mk] theorem sum_mul_sum_eq_sum_perm (hσ : σ.IsCycleOn s) (f g : ι → α) : - ((∑ i ∈ s, f i) * ∑ i ∈ s, g i) = ∑ k ∈ range s.card, ∑ i ∈ s, f i * g ((σ ^ k) i) := + ((∑ i ∈ s, f i) * ∑ i ∈ s, g i) = ∑ k ∈ range #s, ∑ i ∈ s, f i * g ((σ ^ k) i) := sum_smul_sum_eq_sum_perm hσ f g end Finset @@ -1094,10 +1092,10 @@ theorem zpow_eq_ofSubtype_subtypePerm_iff theorem cycle_zpow_mem_support_iff {g : Perm α} (hg : g.IsCycle) {n : ℤ} {x : α} (hx : g x ≠ x) : - (g ^ n) x = x ↔ n % g.support.card = 0 := by - set q := n / g.support.card - set r := n % g.support.card - have div_euc : r + g.support.card * q = n ∧ 0 ≤ r ∧ r < g.support.card := by + (g ^ n) x = x ↔ n % #g.support = 0 := by + set q := n / #g.support + set r := n % #g.support + have div_euc : r + #g.support * q = n ∧ 0 ≤ r ∧ r < #g.support := by rw [← Int.ediv_emod_unique _] · exact ⟨rfl, rfl⟩ simp only [Int.natCast_pos] diff --git a/Mathlib/GroupTheory/Perm/Cycle/Factors.lean b/Mathlib/GroupTheory/Perm/Cycle/Factors.lean index e2fc3224e97be..aa5260cd2a4c3 100644 --- a/Mathlib/GroupTheory/Perm/Cycle/Factors.lean +++ b/Mathlib/GroupTheory/Perm/Cycle/Factors.lean @@ -166,7 +166,7 @@ theorem isCycle_cycleOf (f : Perm α) [DecidableRel f.SameCycle] (hx : f x ≠ x @[simp] theorem two_le_card_support_cycleOf_iff [DecidableEq α] [Fintype α] : - 2 ≤ card (cycleOf f x).support ↔ f x ≠ x := by + 2 ≤ #(cycleOf f x).support ↔ f x ≠ x := by refine ⟨fun h => ?_, fun h => by simpa using (isCycle_cycleOf _ h).two_le_card_support⟩ contrapose! h rw [← cycleOf_eq_one_iff] at h @@ -179,7 +179,7 @@ theorem two_le_card_support_cycleOf_iff [DecidableEq α] [Fintype α] : @[deprecated support_cycleOf_nonempty (since := "2024-06-16")] theorem card_support_cycleOf_pos_iff [DecidableEq α] [Fintype α] : - 0 < card (cycleOf f x).support ↔ f x ≠ x := by + 0 < #(cycleOf f x).support ↔ f x ≠ x := by rw [card_pos, support_cycleOf_nonempty] theorem pow_mod_orderOf_cycleOf_apply (f : Perm α) [DecidableRel f.SameCycle] (n : ℕ) (x : α) : @@ -241,7 +241,7 @@ theorem SameCycle.mem_support_iff {f} [DecidableEq α] [Fintype α] (h : SameCyc support_cycleOf_le f y (mem_support_cycleOf_iff.mpr ⟨h.symm, hy⟩)⟩ theorem pow_mod_card_support_cycleOf_self_apply [DecidableEq α] [Fintype α] - (f : Perm α) (n : ℕ) (x : α) : (f ^ (n % (f.cycleOf x).support.card)) x = (f ^ n) x := by + (f : Perm α) (n : ℕ) (x : α) : (f ^ (n % #(f.cycleOf x).support)) x = (f ^ n) x := by by_cases hx : f x = x · rw [pow_apply_eq_self_of_apply_eq_self hx, pow_apply_eq_self_of_apply_eq_self hx] · rw [← cycleOf_pow_apply_self, ← cycleOf_pow_apply_self f, ← (isCycle_cycleOf f hx).orderOf, @@ -268,17 +268,17 @@ theorem isCycleOn_support_cycleOf [DecidableEq α] [Fintype α] (f : Perm α) (x exact ha.1.symm.trans hb.1⟩ theorem SameCycle.exists_pow_eq_of_mem_support {f} [DecidableEq α] [Fintype α] (h : SameCycle f x y) - (hx : x ∈ f.support) : ∃ i < (f.cycleOf x).support.card, (f ^ i) x = y := by + (hx : x ∈ f.support) : ∃ i < #(f.cycleOf x).support, (f ^ i) x = y := by rw [mem_support] at hx exact Equiv.Perm.IsCycleOn.exists_pow_eq (b := y) (f.isCycleOn_support_cycleOf x) (by rw [mem_support_cycleOf_iff' hx]) (by rwa [mem_support_cycleOf_iff' hx]) theorem SameCycle.exists_pow_eq [DecidableEq α] [Fintype α] (f : Perm α) (h : SameCycle f x y) : - ∃ i : ℕ, 0 < i ∧ i ≤ (f.cycleOf x).support.card + 1 ∧ (f ^ i) x = y := by + ∃ i : ℕ, 0 < i ∧ i ≤ #(f.cycleOf x).support + 1 ∧ (f ^ i) x = y := by by_cases hx : x ∈ f.support · obtain ⟨k, hk, hk'⟩ := h.exists_pow_eq_of_mem_support hx cases' k with k - · refine ⟨(f.cycleOf x).support.card, ?_, self_le_add_right _ _, ?_⟩ + · refine ⟨#(f.cycleOf x).support, ?_, self_le_add_right _ _, ?_⟩ · refine zero_lt_one.trans (one_lt_card_support_of_ne_one ?_) simpa using hx · simp only [pow_zero, coe_one, id_eq] at hk' @@ -293,8 +293,7 @@ theorem SameCycle.exists_pow_eq [DecidableEq α] [Fintype α] (f : Perm α) (h : theorem zpow_eq_zpow_on_iff [DecidableEq α] [Fintype α] (g : Perm α) {m n : ℤ} {x : α} (hx : g x ≠ x) : - (g ^ m) x = (g ^ n) x ↔ - m % (g.cycleOf x).support.card = n % (g.cycleOf x).support.card := by + (g ^ m) x = (g ^ n) x ↔ m % #(g.cycleOf x).support = n % #(g.cycleOf x).support := by rw [Int.emod_eq_emod_iff_emod_sub_eq_zero] conv_lhs => rw [← Int.sub_add_cancel m n, Int.add_comm, zpow_add] simp only [coe_mul, Function.comp_apply, EmbeddingLike.apply_eq_iff_eq] diff --git a/Mathlib/GroupTheory/Perm/Support.lean b/Mathlib/GroupTheory/Perm/Support.lean index 06dbc62cdeb17..10124edc16294 100644 --- a/Mathlib/GroupTheory/Perm/Support.lean +++ b/Mathlib/GroupTheory/Perm/Support.lean @@ -264,8 +264,7 @@ theorem apply_zpow_apply_eq_iff (f : Perm α) (n : ℤ) {x : α} : variable [DecidableEq α] [Fintype α] {f g : Perm α} /-- The `Finset` of nonfixed points of a permutation. -/ -def support (f : Perm α) : Finset α := - univ.filter fun x => f x ≠ x +def support (f : Perm α) : Finset α := {x | f x ≠ x} @[simp] theorem mem_support {x : α} : x ∈ f.support ↔ f x ≠ x := by @@ -526,47 +525,47 @@ theorem support_extend_domain (f : α ≃ Subtype p) {g : Perm α} : exact pb (Subtype.prop _) theorem card_support_extend_domain (f : α ≃ Subtype p) {g : Perm α} : - (g.extendDomain f).support.card = g.support.card := by simp + #(g.extendDomain f).support = #g.support := by simp end ExtendDomain section Card -theorem card_support_eq_zero {f : Perm α} : f.support.card = 0 ↔ f = 1 := by +theorem card_support_eq_zero {f : Perm α} : #f.support = 0 ↔ f = 1 := by rw [Finset.card_eq_zero, support_eq_empty_iff] -theorem one_lt_card_support_of_ne_one {f : Perm α} (h : f ≠ 1) : 1 < f.support.card := by +theorem one_lt_card_support_of_ne_one {f : Perm α} (h : f ≠ 1) : 1 < #f.support := by simp_rw [one_lt_card_iff, mem_support, ← not_or] contrapose! h ext a specialize h (f a) a rwa [apply_eq_iff_eq, or_self_iff, or_self_iff] at h -theorem card_support_ne_one (f : Perm α) : f.support.card ≠ 1 := by +theorem card_support_ne_one (f : Perm α) : #f.support ≠ 1 := by by_cases h : f = 1 · exact ne_of_eq_of_ne (card_support_eq_zero.mpr h) zero_ne_one · exact ne_of_gt (one_lt_card_support_of_ne_one h) @[simp] -theorem card_support_le_one {f : Perm α} : f.support.card ≤ 1 ↔ f = 1 := by +theorem card_support_le_one {f : Perm α} : #f.support ≤ 1 ↔ f = 1 := by rw [le_iff_lt_or_eq, Nat.lt_succ_iff, Nat.le_zero, card_support_eq_zero, or_iff_not_imp_right, imp_iff_right f.card_support_ne_one] -theorem two_le_card_support_of_ne_one {f : Perm α} (h : f ≠ 1) : 2 ≤ f.support.card := +theorem two_le_card_support_of_ne_one {f : Perm α} (h : f ≠ 1) : 2 ≤ #f.support := one_lt_card_support_of_ne_one h theorem card_support_swap_mul {f : Perm α} {x : α} (hx : f x ≠ x) : - (swap x (f x) * f).support.card < f.support.card := + #(swap x (f x) * f).support < #f.support := Finset.card_lt_card ⟨fun _ hz => (mem_support_swap_mul_imp_mem_support_ne hz).left, fun h => absurd (h (mem_support.2 hx)) (mt mem_support.1 (by simp))⟩ -theorem card_support_swap {x y : α} (hxy : x ≠ y) : (swap x y).support.card = 2 := - show (swap x y).support.card = Finset.card ⟨x ::ₘ y ::ₘ 0, by simp [hxy]⟩ from +theorem card_support_swap {x y : α} (hxy : x ≠ y) : #(swap x y).support = 2 := + show #(swap x y).support = #⟨x ::ₘ y ::ₘ 0, by simp [hxy]⟩ from congr_arg card <| by simp [support_swap hxy, *, Finset.ext_iff] @[simp] -theorem card_support_eq_two {f : Perm α} : f.support.card = 2 ↔ IsSwap f := by +theorem card_support_eq_two {f : Perm α} : #f.support = 2 ↔ IsSwap f := by constructor <;> intro h · obtain ⟨x, t, hmem, hins, ht⟩ := card_eq_succ.1 h obtain ⟨y, rfl⟩ := card_eq_one.1 ht @@ -585,7 +584,7 @@ theorem card_support_eq_two {f : Perm α} : f.support.card = 2 ↔ IsSwap f := b exact card_support_swap hxy theorem Disjoint.card_support_mul (h : Disjoint f g) : - (f * g).support.card = f.support.card + g.support.card := by + #(f * g).support = #f.support + #g.support := by rw [← Finset.card_union_of_disjoint] · congr ext @@ -593,7 +592,7 @@ theorem Disjoint.card_support_mul (h : Disjoint f g) : · simpa using h.disjoint_support theorem card_support_prod_list_of_pairwise_disjoint {l : List (Perm α)} (h : l.Pairwise Disjoint) : - l.prod.support.card = (l.map (Finset.card ∘ support)).sum := by + #l.prod.support = (l.map (card ∘ support)).sum := by induction' l with a t ih · exact card_support_eq_zero.mpr rfl · obtain ⟨ha, ht⟩ := List.pairwise_cons.1 h @@ -606,10 +605,8 @@ end support @[simp] theorem support_subtype_perm [DecidableEq α] {s : Finset α} (f : Perm α) (h) : - ((f.subtypePerm h : Perm { x // x ∈ s }).support) = - (s.attach.filter ((fun x => decide (f x ≠ x))) : Finset { x // x ∈ s }) := by - ext - simp [Subtype.ext_iff] + (f.subtypePerm h : Perm s).support = ({x | f x ≠ x} : Finset s) := by + ext; simp [Subtype.ext_iff] end Equiv.Perm @@ -623,7 +620,7 @@ namespace Equiv.Perm variable {α : Type*} theorem fixed_point_card_lt_of_ne_one [DecidableEq α] [Fintype α] {σ : Perm α} (h : σ ≠ 1) : - (filter (fun x => σ x = x) univ).card < Fintype.card α - 1 := by + #{x | σ x = x} < Fintype.card α - 1 := by rw [Nat.lt_sub_iff_add_lt, ← Nat.lt_sub_iff_add_lt', ← Finset.card_compl, Finset.compl_filter] exact one_lt_card_support_of_ne_one h @@ -643,7 +640,7 @@ theorem support_conj : (σ * τ * σ⁻¹).support = τ.support.map σ.toEmbeddi simp only [mem_map_equiv, Perm.coe_mul, Function.comp_apply, Ne, Perm.mem_support, Equiv.eq_symm_apply, inv_def] -theorem card_support_conj : (σ * τ * σ⁻¹).support.card = τ.support.card := by simp +theorem card_support_conj : #(σ * τ * σ⁻¹).support = #τ.support := by simp end Equiv.Perm diff --git a/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean b/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean index 38b5ab4ee0b07..4a709390a2037 100644 --- a/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean +++ b/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean @@ -278,11 +278,11 @@ open scoped Classical @[to_additive IsAddCyclic.card_nsmul_eq_zero_le] theorem IsCyclic.card_pow_eq_one_le [DecidableEq α] [Fintype α] [IsCyclic α] {n : ℕ} (hn0 : 0 < n) : - (univ.filter fun a : α => a ^ n = 1).card ≤ n := + #{a : α | a ^ n = 1} ≤ n := let ⟨g, hg⟩ := IsCyclic.exists_generator (α := α) calc - (univ.filter fun a : α => a ^ n = 1).card ≤ - (zpowers (g ^ (Fintype.card α / Nat.gcd n (Fintype.card α))) : Set α).toFinset.card := + #{a : α | a ^ n = 1} ≤ + #(zpowers (g ^ (Fintype.card α / Nat.gcd n (Fintype.card α))) : Set α).toFinset := card_le_card fun x hx => let ⟨m, hm⟩ := show x ∈ Submonoid.powers g from mem_powers_iff_mem_zpowers.2 <| hg x Set.mem_toFinset.2 @@ -382,21 +382,19 @@ end section Totient -variable [DecidableEq α] [Fintype α] - (hn : ∀ n : ℕ, 0 < n → (univ.filter fun a : α => a ^ n = 1).card ≤ n) +variable [DecidableEq α] [Fintype α] (hn : ∀ n : ℕ, 0 < n → #{a : α | a ^ n = 1} ≤ n) include hn @[to_additive] -private theorem card_pow_eq_one_eq_orderOf_aux (a : α) : - (Finset.univ.filter fun b : α => b ^ orderOf a = 1).card = orderOf a := +private theorem card_pow_eq_one_eq_orderOf_aux (a : α) : #{b : α | b ^ orderOf a = 1} = orderOf a := le_antisymm (hn _ (orderOf_pos a)) (calc orderOf a = @Fintype.card (zpowers a) (id _) := Fintype.card_zpowers.symm _ ≤ - @Fintype.card (↑(univ.filter fun b : α => b ^ orderOf a = 1) : Set α) + @Fintype.card (({b : α | b ^ orderOf a = 1} : Finset _) : Set α) (Fintype.ofFinset _ fun _ => Iff.rfl) := (@Fintype.card_le_of_injective (zpowers a) - (↑(univ.filter fun b : α => b ^ orderOf a = 1) : Set α) (id _) (id _) + (({b : α | b ^ orderOf a = 1} : Finset _) : Set α) (id _) (id _) (fun b => ⟨b.1, mem_filter.2 @@ -405,25 +403,21 @@ private theorem card_pow_eq_one_eq_orderOf_aux (a : α) : rw [← hi, ← zpow_natCast, ← zpow_mul, mul_comm, zpow_mul, zpow_natCast, pow_orderOf_eq_one, one_zpow]⟩⟩) fun _ _ h => Subtype.eq (Subtype.mk.inj h)) - _ = (univ.filter fun b : α => b ^ orderOf a = 1).card := Fintype.card_ofFinset _ _ + _ = #{b : α | b ^ orderOf a = 1} := Fintype.card_ofFinset _ _ ) -- Use φ for `Nat.totient` open Nat @[to_additive] -private theorem card_orderOf_eq_totient_aux₁ : - ∀ {d : ℕ}, - d ∣ Fintype.card α → - 0 < (univ.filter fun a : α => orderOf a = d).card → - (univ.filter fun a : α => orderOf a = d).card = φ d := by - intro d hd hpos +private theorem card_orderOf_eq_totient_aux₁ {d : ℕ} (hd : d ∣ Fintype.card α) + (hpos : 0 < #{a : α | orderOf a = d}) : #{a : α | orderOf a = d} = φ d := by induction' d using Nat.strongRec' with d IH rcases Decidable.eq_or_ne d 0 with (rfl | hd0) · cases Fintype.card_ne_zero (eq_zero_of_zero_dvd hd) - rcases card_pos.1 hpos with ⟨a, ha'⟩ + rcases Finset.card_pos.1 hpos with ⟨a, ha'⟩ have ha : orderOf a = d := (mem_filter.1 ha').2 have h1 : - (∑ m ∈ d.properDivisors, (univ.filter fun a : α => orderOf a = m).card) = + (∑ m ∈ d.properDivisors, #{a : α | orderOf a = m}) = ∑ m ∈ d.properDivisors, φ m := by refine Finset.sum_congr rfl fun m hm => ?_ simp only [mem_filter, mem_range, mem_properDivisors] at hm @@ -431,7 +425,7 @@ private theorem card_orderOf_eq_totient_aux₁ : simp only [mem_filter, mem_univ, orderOf_pow a, ha, true_and, Nat.gcd_eq_right (div_dvd_of_dvd hm.1), Nat.div_div_self hm.1 hd0] have h2 : - (∑ m ∈ d.divisors, (univ.filter fun a : α => orderOf a = m).card) = + (∑ m ∈ d.divisors, #{a : α | orderOf a = m}) = ∑ m ∈ d.divisors, φ m := by rw [← filter_dvd_eq_divisors hd0, sum_card_orderOf_eq_card_pow_eq_one hd0, filter_dvd_eq_divisors hd0, sum_totient, ← ha, card_pow_eq_one_eq_orderOf_aux hn a] @@ -439,7 +433,7 @@ private theorem card_orderOf_eq_totient_aux₁ : @[to_additive] theorem card_orderOf_eq_totient_aux₂ {d : ℕ} (hd : d ∣ Fintype.card α) : - (univ.filter fun a : α => orderOf a = d).card = φ d := by + #{a : α | orderOf a = d} = φ d := by let c := Fintype.card α have hc0 : 0 < c := Fintype.card_pos_iff.2 ⟨1⟩ apply card_orderOf_eq_totient_aux₁ hn hd @@ -448,11 +442,11 @@ theorem card_orderOf_eq_totient_aux₂ {d : ℕ} (hd : d ∣ Fintype.card α) : simp_rw [not_lt, Nat.le_zero, Finset.card_eq_zero] at h0 apply lt_irrefl c calc - c = ∑ m ∈ c.divisors, (univ.filter fun a : α => orderOf a = m).card := by + c = ∑ m ∈ c.divisors, #{a : α | orderOf a = m} := by simp only [← filter_dvd_eq_divisors hc0.ne', sum_card_orderOf_eq_card_pow_eq_one hc0.ne'] apply congr_arg card simp [c] - _ = ∑ m ∈ c.divisors.erase d, (univ.filter fun a : α => orderOf a = m).card := by + _ = ∑ m ∈ c.divisors.erase d, #{a : α | orderOf a = m} := by rw [eq_comm] refine sum_subset (erase_subset _ _) fun m hm₁ hm₂ => ?_ have : m = d := by @@ -464,7 +458,7 @@ theorem card_orderOf_eq_totient_aux₂ {d : ℕ} (hd : d ∣ Fintype.card α) : have hmc : m ∣ c := by simp only [mem_erase, mem_divisors] at hm tauto - rcases (filter (fun a : α => orderOf a = m) univ).card.eq_zero_or_pos with (h1 | h1) + obtain h1 | h1 := (#{a : α | orderOf a = m}).eq_zero_or_pos · simp [h1] · simp [card_orderOf_eq_totient_aux₁ hn hmc h1] _ < ∑ m ∈ c.divisors, φ m := @@ -473,7 +467,7 @@ theorem card_orderOf_eq_totient_aux₂ {d : ℕ} (hd : d ∣ Fintype.card α) : @[to_additive isAddCyclic_of_card_nsmul_eq_zero_le] theorem isCyclic_of_card_pow_eq_one_le : IsCyclic α := - have : (univ.filter fun a : α => orderOf a = Fintype.card α).Nonempty := + have : Finset.Nonempty {a : α | orderOf a = Fintype.card α} := card_pos.1 <| by rw [card_orderOf_eq_totient_aux₂ hn dvd_rfl, totient_pos] apply Fintype.card_pos @@ -486,8 +480,8 @@ alias isAddCyclic_of_card_pow_eq_one_le := isAddCyclic_of_card_nsmul_eq_zero_le end Totient @[to_additive] -theorem IsCyclic.card_orderOf_eq_totient [IsCyclic α] [Fintype α] {d : ℕ} - (hd : d ∣ Fintype.card α) : (univ.filter fun a : α => orderOf a = d).card = totient d := by +lemma IsCyclic.card_orderOf_eq_totient [IsCyclic α] [Fintype α] {d : ℕ} (hd : d ∣ Fintype.card α) : + #{a : α | orderOf a = d} = totient d := by classical apply card_orderOf_eq_totient_aux₂ (fun n => IsCyclic.card_pow_eq_one_le) hd @[deprecated (since := "2024-02-21")] diff --git a/Mathlib/InformationTheory/Hamming.lean b/Mathlib/InformationTheory/Hamming.lean index 1cd0c586d3812..38a74feba3a9d 100644 --- a/Mathlib/InformationTheory/Hamming.lean +++ b/Mathlib/InformationTheory/Hamming.lean @@ -34,8 +34,7 @@ variable {α ι : Type*} {β : ι → Type*} [Fintype ι] [∀ i, DecidableEq ( variable {γ : ι → Type*} [∀ i, DecidableEq (γ i)] /-- The Hamming distance function to the naturals. -/ -def hammingDist (x y : ∀ i, β i) : ℕ := - (univ.filter fun i => x i ≠ y i).card +def hammingDist (x y : ∀ i, β i) : ℕ := #{i | x i ≠ y i} /-- Corresponds to `dist_self`. -/ @[simp] @@ -132,8 +131,7 @@ section Zero variable [∀ i, Zero (β i)] [∀ i, Zero (γ i)] /-- The Hamming weight function to the naturals. -/ -def hammingNorm (x : ∀ i, β i) : ℕ := - (univ.filter (x · ≠ 0)).card +def hammingNorm (x : ∀ i, β i) : ℕ := #{i | x i ≠ 0} /-- Corresponds to `dist_zero_right`. -/ @[simp] diff --git a/Mathlib/LinearAlgebra/AffineSpace/Combination.lean b/Mathlib/LinearAlgebra/AffineSpace/Combination.lean index c3adfc82235fd..33d82512d88eb 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/Combination.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/Combination.lean @@ -202,18 +202,18 @@ theorem weightedVSubOfPoint_sdiff_sub [DecidableEq ι] {s₂ : Finset ι} (h : s s.weightedVSubOfPoint p b w := by rw [map_neg, sub_neg_eq_add, s.weightedVSubOfPoint_sdiff h] -/-- A weighted sum over `s.subtype pred` equals one over `s.filter pred`. -/ +/-- A weighted sum over `s.subtype pred` equals one over `{x ∈ s | pred x}`. -/ theorem weightedVSubOfPoint_subtype_eq_filter (w : ι → k) (p : ι → P) (b : P) (pred : ι → Prop) [DecidablePred pred] : ((s.subtype pred).weightedVSubOfPoint (fun i => p i) b fun i => w i) = - (s.filter pred).weightedVSubOfPoint p b w := by + {x ∈ s | pred x}.weightedVSubOfPoint p b w := by rw [weightedVSubOfPoint_apply, weightedVSubOfPoint_apply, ← sum_subtype_eq_sum_filter] -/-- A weighted sum over `s.filter pred` equals one over `s` if all the weights at indices in `s` +/-- A weighted sum over `{x ∈ s | pred x}` equals one over `s` if all the weights at indices in `s` not satisfying `pred` are zero. -/ theorem weightedVSubOfPoint_filter_of_ne (w : ι → k) (p : ι → P) (b : P) {pred : ι → Prop} [DecidablePred pred] (h : ∀ i ∈ s, w i ≠ 0 → pred i) : - (s.filter pred).weightedVSubOfPoint p b w = s.weightedVSubOfPoint p b w := by + {x ∈ s | pred x}.weightedVSubOfPoint p b w = s.weightedVSubOfPoint p b w := by rw [weightedVSubOfPoint_apply, weightedVSubOfPoint_apply, sum_filter_of_ne] intro i hi hne refine h i hi ?_ @@ -319,17 +319,17 @@ theorem weightedVSub_sdiff_sub [DecidableEq ι] {s₂ : Finset ι} (h : s₂ ⊆ (p : ι → P) : (s \ s₂).weightedVSub p w - s₂.weightedVSub p (-w) = s.weightedVSub p w := s.weightedVSubOfPoint_sdiff_sub h _ _ _ -/-- A weighted sum over `s.subtype pred` equals one over `s.filter pred`. -/ +/-- A weighted sum over `s.subtype pred` equals one over `{x ∈ s | pred x}`. -/ theorem weightedVSub_subtype_eq_filter (w : ι → k) (p : ι → P) (pred : ι → Prop) [DecidablePred pred] : ((s.subtype pred).weightedVSub (fun i => p i) fun i => w i) = - (s.filter pred).weightedVSub p w := + {x ∈ s | pred x}.weightedVSub p w := s.weightedVSubOfPoint_subtype_eq_filter _ _ _ _ -/-- A weighted sum over `s.filter pred` equals one over `s` if all the weights at indices in `s` +/-- A weighted sum over `{x ∈ s | pred x}` equals one over `s` if all the weights at indices in `s` not satisfying `pred` are zero. -/ theorem weightedVSub_filter_of_ne (w : ι → k) (p : ι → P) {pred : ι → Prop} [DecidablePred pred] - (h : ∀ i ∈ s, w i ≠ 0 → pred i) : (s.filter pred).weightedVSub p w = s.weightedVSub p w := + (h : ∀ i ∈ s, w i ≠ 0 → pred i) : {x ∈ s | pred x}.weightedVSub p w = s.weightedVSub p w := s.weightedVSubOfPoint_filter_of_ne _ _ _ h /-- A constant multiplier of the weights in `weightedVSub_of` may be moved outside the sum. -/ @@ -497,7 +497,7 @@ theorem affineCombination_sdiff_sub [DecidableEq ι] {s₂ : Finset ι} (h : s the affine combination of the other points with the given weights. -/ theorem affineCombination_eq_of_weightedVSub_eq_zero_of_eq_neg_one {w : ι → k} {p : ι → P} (hw : s.weightedVSub p w = (0 : V)) {i : ι} [DecidablePred (· ≠ i)] (his : i ∈ s) - (hwi : w i = -1) : (s.filter (· ≠ i)).affineCombination k p w = p i := by + (hwi : w i = -1) : {x ∈ s | x ≠ i}.affineCombination k p w = p i := by classical rw [← @vsub_eq_zero_iff_eq V, ← hw, ← s.affineCombination_sdiff_sub (singleton_subset_iff.2 his), sdiff_singleton_eq_erase, @@ -507,18 +507,18 @@ theorem affineCombination_eq_of_weightedVSub_eq_zero_of_eq_neg_one {w : ι → k · simp [hwi] · simp -/-- An affine combination over `s.subtype pred` equals one over `s.filter pred`. -/ +/-- An affine combination over `s.subtype pred` equals one over `{x ∈ s | pred x}`. -/ theorem affineCombination_subtype_eq_filter (w : ι → k) (p : ι → P) (pred : ι → Prop) [DecidablePred pred] : ((s.subtype pred).affineCombination k (fun i => p i) fun i => w i) = - (s.filter pred).affineCombination k p w := by + {x ∈ s | pred x}.affineCombination k p w := by rw [affineCombination_apply, affineCombination_apply, weightedVSubOfPoint_subtype_eq_filter] -/-- An affine combination over `s.filter pred` equals one over `s` if all the weights at indices +/-- An affine combination over `{x ∈ s | pred x}` equals one over `s` if all the weights at indices in `s` not satisfying `pred` are zero. -/ theorem affineCombination_filter_of_ne (w : ι → k) (p : ι → P) {pred : ι → Prop} [DecidablePred pred] (h : ∀ i ∈ s, w i ≠ 0 → pred i) : - (s.filter pred).affineCombination k p w = s.affineCombination k p w := by + {x ∈ s | pred x}.affineCombination k p w = s.affineCombination k p w := by rw [affineCombination_apply, affineCombination_apply, s.weightedVSubOfPoint_filter_of_ne _ _ _ h] @@ -712,29 +712,29 @@ variable [AffineSpace V P] {ι : Type*} (s : Finset ι) {ι₂ : Type*} (s₂ : /-- The weights for the centroid of some points. -/ def centroidWeights : ι → k := - Function.const ι (card s : k)⁻¹ + Function.const ι (#s : k)⁻¹ /-- `centroidWeights` at any point. -/ @[simp] -theorem centroidWeights_apply (i : ι) : s.centroidWeights k i = (card s : k)⁻¹ := +theorem centroidWeights_apply (i : ι) : s.centroidWeights k i = (#s : k)⁻¹ := rfl /-- `centroidWeights` equals a constant function. -/ -theorem centroidWeights_eq_const : s.centroidWeights k = Function.const ι (card s : k)⁻¹ := +theorem centroidWeights_eq_const : s.centroidWeights k = Function.const ι (#s : k)⁻¹ := rfl variable {k} /-- The weights in the centroid sum to 1, if the number of points, converted to `k`, is not zero. -/ -theorem sum_centroidWeights_eq_one_of_cast_card_ne_zero (h : (card s : k) ≠ 0) : +theorem sum_centroidWeights_eq_one_of_cast_card_ne_zero (h : (#s : k) ≠ 0) : ∑ i ∈ s, s.centroidWeights k i = 1 := by simp [h] variable (k) /-- In the characteristic zero case, the weights in the centroid sum to 1 if the number of points is not zero. -/ -theorem sum_centroidWeights_eq_one_of_card_ne_zero [CharZero k] (h : card s ≠ 0) : +theorem sum_centroidWeights_eq_one_of_card_ne_zero [CharZero k] (h : #s ≠ 0) : ∑ i ∈ s, s.centroidWeights k i = 1 := by -- Porting note: `simp` cannot find `mul_inv_cancel` and does not use `norm_cast` simp only [centroidWeights_apply, sum_const, nsmul_eq_mul, ne_eq, Nat.cast_eq_zero, card_eq_zero] @@ -749,7 +749,7 @@ theorem sum_centroidWeights_eq_one_of_nonempty [CharZero k] (h : s.Nonempty) : /-- In the characteristic zero case, the weights in the centroid sum to 1 if the number of points is `n + 1`. -/ -theorem sum_centroidWeights_eq_one_of_card_eq_add_one [CharZero k] {n : ℕ} (h : card s = n + 1) : +theorem sum_centroidWeights_eq_one_of_card_eq_add_one [CharZero k] {n : ℕ} (h : #s = n + 1) : ∑ i ∈ s, s.centroidWeights k i = 1 := s.sum_centroidWeights_eq_one_of_card_ne_zero k (h.symm ▸ Nat.succ_ne_zero n) @@ -780,7 +780,7 @@ theorem centroid_pair [DecidableEq ι] [Invertible (2 : k)] (p : ι → P) (i₁ ({i₁, i₂} : Finset ι).centroid k p = (2⁻¹ : k) • (p i₂ -ᵥ p i₁) +ᵥ p i₁ := by by_cases h : i₁ = i₂ · simp [h] - · have hc : (card ({i₁, i₂} : Finset ι) : k) ≠ 0 := by + · have hc : (#{i₁, i₂} : k) ≠ 0 := by rw [card_insert_of_not_mem (not_mem_singleton.2 h), card_singleton] norm_num exact Invertible.ne_zero _ @@ -825,7 +825,7 @@ theorem sum_centroidWeightsIndicator [Fintype ι] : indexed by a `Fintype` sum to 1 if the number of points is not zero. -/ theorem sum_centroidWeightsIndicator_eq_one_of_card_ne_zero [CharZero k] [Fintype ι] - (h : card s ≠ 0) : ∑ i, s.centroidWeightsIndicator k i = 1 := by + (h : #s ≠ 0) : ∑ i, s.centroidWeightsIndicator k i = 1 := by rw [sum_centroidWeightsIndicator] exact s.sum_centroidWeights_eq_one_of_card_ne_zero k h @@ -839,7 +839,7 @@ theorem sum_centroidWeightsIndicator_eq_one_of_nonempty [CharZero k] [Fintype ι /-- In the characteristic zero case, the weights in the centroid indexed by a `Fintype` sum to 1 if the number of points is `n + 1`. -/ theorem sum_centroidWeightsIndicator_eq_one_of_card_eq_add_one [CharZero k] [Fintype ι] {n : ℕ} - (h : card s = n + 1) : ∑ i, s.centroidWeightsIndicator k i = 1 := by + (h : #s = n + 1) : ∑ i, s.centroidWeightsIndicator k i = 1 := by rw [sum_centroidWeightsIndicator] exact s.sum_centroidWeights_eq_one_of_card_eq_add_one k h @@ -1099,7 +1099,7 @@ open Set Finset /-- The centroid lies in the affine span if the number of points, converted to `k`, is not zero. -/ theorem centroid_mem_affineSpan_of_cast_card_ne_zero {s : Finset ι} (p : ι → P) - (h : (card s : k) ≠ 0) : s.centroid k p ∈ affineSpan k (range p) := + (h : (#s : k) ≠ 0) : s.centroid k p ∈ affineSpan k (range p) := affineCombination_mem_affineSpan (s.sum_centroidWeights_eq_one_of_cast_card_ne_zero h) p variable (k) @@ -1107,7 +1107,7 @@ variable (k) /-- In the characteristic zero case, the centroid lies in the affine span if the number of points is not zero. -/ theorem centroid_mem_affineSpan_of_card_ne_zero [CharZero k] {s : Finset ι} (p : ι → P) - (h : card s ≠ 0) : s.centroid k p ∈ affineSpan k (range p) := + (h : #s ≠ 0) : s.centroid k p ∈ affineSpan k (range p) := affineCombination_mem_affineSpan (s.sum_centroidWeights_eq_one_of_card_ne_zero k h) p /-- In the characteristic zero case, the centroid lies in the affine @@ -1119,7 +1119,7 @@ theorem centroid_mem_affineSpan_of_nonempty [CharZero k] {s : Finset ι} (p : ι /-- In the characteristic zero case, the centroid lies in the affine span if the number of points is `n + 1`. -/ theorem centroid_mem_affineSpan_of_card_eq_add_one [CharZero k] {s : Finset ι} (p : ι → P) {n : ℕ} - (h : card s = n + 1) : s.centroid k p ∈ affineSpan k (range p) := + (h : #s = n + 1) : s.centroid k p ∈ affineSpan k (range p) := affineCombination_mem_affineSpan (s.sum_centroidWeights_eq_one_of_card_eq_add_one k h) p end DivisionRing diff --git a/Mathlib/LinearAlgebra/AffineSpace/FiniteDimensional.lean b/Mathlib/LinearAlgebra/AffineSpace/FiniteDimensional.lean index 935dc82abfb3b..0ef068c764695 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/FiniteDimensional.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/FiniteDimensional.lean @@ -23,6 +23,7 @@ subspaces of affine spaces. noncomputable section open Affine +open scoped Finset section AffineSpace' @@ -87,18 +88,18 @@ variable {k} /-- The `vectorSpan` of a finite subset of an affinely independent family has dimension one less than its cardinality. -/ theorem AffineIndependent.finrank_vectorSpan_image_finset [DecidableEq P] - {p : ι → P} (hi : AffineIndependent k p) {s : Finset ι} {n : ℕ} (hc : Finset.card s = n + 1) : + {p : ι → P} (hi : AffineIndependent k p) {s : Finset ι} {n : ℕ} (hc : #s = n + 1) : finrank k (vectorSpan k (s.image p : Set P)) = n := by classical have hi' := hi.range.mono (Set.image_subset_range p ↑s) - have hc' : (s.image p).card = n + 1 := by rwa [s.card_image_of_injective hi.injective] + have hc' : #(s.image p) = n + 1 := by rwa [s.card_image_of_injective hi.injective] have hn : (s.image p).Nonempty := by simp [hc', ← Finset.card_pos] rcases hn with ⟨p₁, hp₁⟩ have hp₁' : p₁ ∈ p '' s := by simpa using hp₁ rw [affineIndependent_set_iff_linearIndependent_vsub k hp₁', ← Finset.coe_singleton, ← Finset.coe_image, ← Finset.coe_sdiff, Finset.sdiff_singleton_eq_erase, ← Finset.coe_image] at hi' - have hc : (Finset.image (fun p : P => p -ᵥ p₁) ((Finset.image p s).erase p₁)).card = n := by + have hc : #(((s.image p).erase p₁).image (· -ᵥ p₁)) = n := by rw [Finset.card_image_of_injective _ (vsub_left_injective _), Finset.card_erase_of_mem hp₁] exact Nat.pred_eq_of_eq_succ hc' rwa [vectorSpan_eq_span_vsub_finset_right_ne k hp₁, finrank_span_finset_eq_card, hc] @@ -132,7 +133,7 @@ variable (k) /-- The `vectorSpan` of `n + 1` points in an indexed family has dimension at most `n`. -/ theorem finrank_vectorSpan_image_finset_le [DecidableEq P] (p : ι → P) (s : Finset ι) {n : ℕ} - (hc : Finset.card s = n + 1) : finrank k (vectorSpan k (s.image p : Set P)) ≤ n := by + (hc : #s = n + 1) : finrank k (vectorSpan k (s.image p : Set P)) ≤ n := by classical have hn : (s.image p).Nonempty := by rw [Finset.image_nonempty, ← Finset.card_pos, hc] @@ -214,7 +215,7 @@ open Finset in cardinality is at most the cardinality of that finset. -/ lemma AffineIndependent.card_le_card_of_subset_affineSpan {s t : Finset V} (hs : AffineIndependent k ((↑) : s → V)) (hst : (s : Set V) ⊆ affineSpan k (t : Set V)) : - s.card ≤ t.card := by + #s ≤ #t := by obtain rfl | hs' := s.eq_empty_or_nonempty · simp obtain rfl | ht' := t.eq_empty_or_nonempty @@ -234,7 +235,7 @@ open Finset in another finset, then its cardinality is strictly less than the cardinality of that finset. -/ lemma AffineIndependent.card_lt_card_of_affineSpan_lt_affineSpan {s t : Finset V} (hs : AffineIndependent k ((↑) : s → V)) - (hst : affineSpan k (s : Set V) < affineSpan k (t : Set V)) : s.card < t.card := by + (hst : affineSpan k (s : Set V) < affineSpan k (t : Set V)) : #s < #t := by obtain rfl | hs' := s.eq_empty_or_nonempty · simpa [card_pos] using hst obtain rfl | ht' := t.eq_empty_or_nonempty @@ -255,7 +256,7 @@ cardinality, it equals that submodule. -/ theorem AffineIndependent.vectorSpan_image_finset_eq_of_le_of_card_eq_finrank_add_one [DecidableEq P] {p : ι → P} (hi : AffineIndependent k p) {s : Finset ι} {sm : Submodule k V} [FiniteDimensional k sm] - (hle : vectorSpan k (s.image p : Set P) ≤ sm) (hc : Finset.card s = finrank k sm + 1) : + (hle : vectorSpan k (s.image p : Set P) ≤ sm) (hc : #s = finrank k sm + 1) : vectorSpan k (s.image p : Set P) = sm := Submodule.eq_of_le_of_finrank_eq hle <| hi.finrank_vectorSpan_image_finset hc @@ -275,7 +276,7 @@ theorem AffineIndependent.affineSpan_image_finset_eq_of_le_of_card_eq_finrank_ad [DecidableEq P] {p : ι → P} (hi : AffineIndependent k p) {s : Finset ι} {sp : AffineSubspace k P} [FiniteDimensional k sp.direction] (hle : affineSpan k (s.image p : Set P) ≤ sp) - (hc : Finset.card s = finrank k sp.direction + 1) : affineSpan k (s.image p : Set P) = sp := by + (hc : #s = finrank k sp.direction + 1) : affineSpan k (s.image p : Set P) = sp := by have hn : s.Nonempty := by rw [← Finset.card_pos, hc] apply Nat.succ_pos diff --git a/Mathlib/LinearAlgebra/AffineSpace/Independent.lean b/Mathlib/LinearAlgebra/AffineSpace/Independent.lean index 0fe67bc39f3f9..f8801772f8d02 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/Independent.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/Independent.lean @@ -810,19 +810,19 @@ add_decl_doc Affine.Simplex.ext_iff /-- A face of a simplex is a simplex with the given subset of points. -/ -def face {n : ℕ} (s : Simplex k P n) {fs : Finset (Fin (n + 1))} {m : ℕ} (h : fs.card = m + 1) : +def face {n : ℕ} (s : Simplex k P n) {fs : Finset (Fin (n + 1))} {m : ℕ} (h : #fs = m + 1) : Simplex k P m := ⟨s.points ∘ fs.orderEmbOfFin h, s.independent.comp_embedding (fs.orderEmbOfFin h).toEmbedding⟩ /-- The points of a face of a simplex are given by `mono_of_fin`. -/ theorem face_points {n : ℕ} (s : Simplex k P n) {fs : Finset (Fin (n + 1))} {m : ℕ} - (h : fs.card = m + 1) (i : Fin (m + 1)) : + (h : #fs = m + 1) (i : Fin (m + 1)) : (s.face h).points i = s.points (fs.orderEmbOfFin h i) := rfl /-- The points of a face of a simplex are given by `mono_of_fin`. -/ theorem face_points' {n : ℕ} (s : Simplex k P n) {fs : Finset (Fin (n + 1))} {m : ℕ} - (h : fs.card = m + 1) : (s.face h).points = s.points ∘ fs.orderEmbOfFin h := + (h : #fs = m + 1) : (s.face h).points = s.points ∘ fs.orderEmbOfFin h := rfl /-- A single-point face equals the 0-simplex constructed with @@ -838,7 +838,7 @@ theorem face_eq_mkOfPoint {n : ℕ} (s : Simplex k P n) (i : Fin (n + 1)) : /-- The set of points of a face. -/ @[simp] theorem range_face_points {n : ℕ} (s : Simplex k P n) {fs : Finset (Fin (n + 1))} {m : ℕ} - (h : fs.card = m + 1) : Set.range (s.face h).points = s.points '' ↑fs := by + (h : #fs = m + 1) : Set.range (s.face h).points = s.points '' ↑fs := by rw [face_points', Set.range_comp, Finset.range_orderEmbOfFin] /-- Remap a simplex along an `Equiv` of index types. -/ @@ -889,7 +889,7 @@ variable {k : Type*} {V : Type*} {P : Type*} [DivisionRing k] [AddCommGroup V] [ the points. -/ @[simp] theorem face_centroid_eq_centroid {n : ℕ} (s : Simplex k P n) {fs : Finset (Fin (n + 1))} {m : ℕ} - (h : fs.card = m + 1) : Finset.univ.centroid k (s.face h).points = fs.centroid k s.points := by + (h : #fs = m + 1) : Finset.univ.centroid k (s.face h).points = fs.centroid k s.points := by convert (Finset.univ.centroid_map k (fs.orderEmbOfFin h).toEmbedding s.points).symm rw [← Finset.coe_inj, Finset.coe_map, Finset.coe_univ, Set.image_univ] simp @@ -899,7 +899,7 @@ two subsets of the points of a simplex are equal if and only if those faces are given by the same subset of points. -/ @[simp] theorem centroid_eq_iff [CharZero k] {n : ℕ} (s : Simplex k P n) {fs₁ fs₂ : Finset (Fin (n + 1))} - {m₁ m₂ : ℕ} (h₁ : fs₁.card = m₁ + 1) (h₂ : fs₂.card = m₂ + 1) : + {m₁ m₂ : ℕ} (h₁ : #fs₁ = m₁ + 1) (h₂ : #fs₂ = m₂ + 1) : fs₁.centroid k s.points = fs₂.centroid k s.points ↔ fs₁ = fs₂ := by refine ⟨fun h => ?_, @congrArg _ _ fs₁ fs₂ (fun z => Finset.centroid k z s.points)⟩ rw [Finset.centroid_eq_affineCombination_fintype, @@ -924,7 +924,7 @@ theorem centroid_eq_iff [CharZero k] {n : ℕ} (s : Simplex k P n) {fs₁ fs₂ faces of a simplex are equal if and only if those faces are given by the same subset of points. -/ theorem face_centroid_eq_iff [CharZero k] {n : ℕ} (s : Simplex k P n) - {fs₁ fs₂ : Finset (Fin (n + 1))} {m₁ m₂ : ℕ} (h₁ : fs₁.card = m₁ + 1) (h₂ : fs₂.card = m₂ + 1) : + {fs₁ fs₂ : Finset (Fin (n + 1))} {m₁ m₂ : ℕ} (h₁ : #fs₁ = m₁ + 1) (h₂ : #fs₂ = m₂ + 1) : Finset.univ.centroid k (s.face h₁).points = Finset.univ.centroid k (s.face h₂).points ↔ fs₁ = fs₂ := by rw [face_centroid_eq_centroid, face_centroid_eq_centroid] diff --git a/Mathlib/LinearAlgebra/Lagrange.lean b/Mathlib/LinearAlgebra/Lagrange.lean index 9b87bad3dd5fa..b683703a0116d 100644 --- a/Mathlib/LinearAlgebra/Lagrange.lean +++ b/Mathlib/LinearAlgebra/Lagrange.lean @@ -36,10 +36,11 @@ variable {R : Type*} [CommRing R] [IsDomain R] {f g : R[X]} section Finset open Function Fintype +open scoped Finset variable (s : Finset R) -theorem eq_zero_of_degree_lt_of_eval_finset_eq_zero (degree_f_lt : f.degree < s.card) +theorem eq_zero_of_degree_lt_of_eval_finset_eq_zero (degree_f_lt : f.degree < #s) (eval_f : ∀ x ∈ s, f.eval x = 0) : f = 0 := by rw [← mem_degreeLT] at degree_f_lt simp_rw [eval_eq_sum_degreeLTEquiv degree_f_lt] at eval_f @@ -49,15 +50,15 @@ theorem eq_zero_of_degree_lt_of_eval_finset_eq_zero (degree_f_lt : f.degree < s. (Injective.comp (Embedding.subtype _).inj' (equivFinOfCardEq (card_coe _)).symm.injective) fun _ => eval_f _ (Finset.coe_mem _) -theorem eq_of_degree_sub_lt_of_eval_finset_eq (degree_fg_lt : (f - g).degree < s.card) +theorem eq_of_degree_sub_lt_of_eval_finset_eq (degree_fg_lt : (f - g).degree < #s) (eval_fg : ∀ x ∈ s, f.eval x = g.eval x) : f = g := by rw [← sub_eq_zero] refine eq_zero_of_degree_lt_of_eval_finset_eq_zero _ degree_fg_lt ?_ simp_rw [eval_sub, sub_eq_zero] exact eval_fg -theorem eq_of_degrees_lt_of_eval_finset_eq (degree_f_lt : f.degree < s.card) - (degree_g_lt : g.degree < s.card) (eval_fg : ∀ x ∈ s, f.eval x = g.eval x) : f = g := by +theorem eq_of_degrees_lt_of_eval_finset_eq (degree_f_lt : f.degree < #s) + (degree_g_lt : g.degree < #s) (eval_fg : ∀ x ∈ s, f.eval x = g.eval x) : f = g := by rw [← mem_degreeLT] at degree_f_lt degree_g_lt refine eq_of_degree_sub_lt_of_eval_finset_eq _ ?_ eval_fg rw [← mem_degreeLT]; exact Submodule.sub_mem _ degree_f_lt degree_g_lt @@ -67,7 +68,7 @@ Two polynomials, with the same degree and leading coefficient, which have the sa on a set of distinct values with cardinality equal to the degree, are equal. -/ theorem eq_of_degree_le_of_eval_finset_eq - (h_deg_le : f.degree ≤ s.card) + (h_deg_le : f.degree ≤ #s) (h_deg_eq : f.degree = g.degree) (hlc : f.leadingCoeff = g.leadingCoeff) (h_eval : ∀ x ∈ s, f.eval x = g.eval x) : @@ -86,7 +87,7 @@ open Finset variable {ι : Type*} {v : ι → R} (s : Finset ι) theorem eq_zero_of_degree_lt_of_eval_index_eq_zero (hvs : Set.InjOn v s) - (degree_f_lt : f.degree < s.card) (eval_f : ∀ i ∈ s, f.eval (v i) = 0) : f = 0 := by + (degree_f_lt : f.degree < #s) (eval_f : ∀ i ∈ s, f.eval (v i) = 0) : f = 0 := by classical rw [← card_image_of_injOn hvs] at degree_f_lt refine eq_zero_of_degree_lt_of_eval_finset_eq_zero _ degree_f_lt ?_ @@ -95,21 +96,21 @@ theorem eq_zero_of_degree_lt_of_eval_index_eq_zero (hvs : Set.InjOn v s) exact eval_f _ hj theorem eq_of_degree_sub_lt_of_eval_index_eq (hvs : Set.InjOn v s) - (degree_fg_lt : (f - g).degree < s.card) (eval_fg : ∀ i ∈ s, f.eval (v i) = g.eval (v i)) : + (degree_fg_lt : (f - g).degree < #s) (eval_fg : ∀ i ∈ s, f.eval (v i) = g.eval (v i)) : f = g := by rw [← sub_eq_zero] refine eq_zero_of_degree_lt_of_eval_index_eq_zero _ hvs degree_fg_lt ?_ simp_rw [eval_sub, sub_eq_zero] exact eval_fg -theorem eq_of_degrees_lt_of_eval_index_eq (hvs : Set.InjOn v s) (degree_f_lt : f.degree < s.card) - (degree_g_lt : g.degree < s.card) (eval_fg : ∀ i ∈ s, f.eval (v i) = g.eval (v i)) : f = g := by +theorem eq_of_degrees_lt_of_eval_index_eq (hvs : Set.InjOn v s) (degree_f_lt : f.degree < #s) + (degree_g_lt : g.degree < #s) (eval_fg : ∀ i ∈ s, f.eval (v i) = g.eval (v i)) : f = g := by refine eq_of_degree_sub_lt_of_eval_index_eq _ hvs ?_ eval_fg rw [← mem_degreeLT] at degree_f_lt degree_g_lt ⊢ exact Submodule.sub_mem _ degree_f_lt degree_g_lt theorem eq_of_degree_le_of_eval_index_eq (hvs : Set.InjOn v s) - (h_deg_le : f.degree ≤ s.card) + (h_deg_le : f.degree ≤ #s) (h_deg_eq : f.degree = g.degree) (hlc : f.leadingCoeff = g.leadingCoeff) (h_eval : ∀ i ∈ s, f.eval (v i) = g.eval (v i)) : f = g := by @@ -233,7 +234,7 @@ theorem eval_basis_of_ne (hij : i ≠ j) (hj : j ∈ s) : (Lagrange.basis s v i) @[simp] theorem natDegree_basis (hvs : Set.InjOn v s) (hi : i ∈ s) : - (Lagrange.basis s v i).natDegree = s.card - 1 := by + (Lagrange.basis s v i).natDegree = #s - 1 := by have H : ∀ j, j ∈ s.erase i → basisDivisor (v i) (v j) ≠ 0 := by simp_rw [Ne, mem_erase, basisDivisor_eq_zero_iff] exact fun j ⟨hij₁, hj⟩ hij₂ => hij₁ (hvs hj hi hij₂.symm) @@ -244,14 +245,14 @@ theorem natDegree_basis (hvs : Set.InjOn v s) (hi : i ∈ s) : exact H _ hj theorem degree_basis (hvs : Set.InjOn v s) (hi : i ∈ s) : - (Lagrange.basis s v i).degree = ↑(s.card - 1) := by + (Lagrange.basis s v i).degree = ↑(#s - 1) := by rw [degree_eq_natDegree (basis_ne_zero hvs hi), natDegree_basis hvs hi] -- Porting note: Added `Nat.cast_withBot` rewrites theorem sum_basis (hvs : Set.InjOn v s) (hs : s.Nonempty) : ∑ j ∈ s, Lagrange.basis s v j = 1 := by refine eq_of_degrees_lt_of_eval_index_eq s hvs (lt_of_le_of_lt (degree_sum_le _ _) ?_) ?_ ?_ - · rw [Nat.cast_withBot, Finset.sup_lt_iff (WithBot.bot_lt_coe s.card)] + · rw [Nat.cast_withBot, Finset.sup_lt_iff (WithBot.bot_lt_coe #s)] intro i hi rw [degree_basis hvs hi, Nat.cast_withBot, WithBot.coe_lt_coe] exact Nat.pred_lt (card_ne_zero_of_mem hi) @@ -282,7 +283,7 @@ open Finset /-- Lagrange interpolation: given a finset `s : Finset ι`, a nodal map `v : ι → F` injective on `s` and a value function `r : ι → F`, `interpolate s v r` is the unique -polynomial of degree `< s.card` that takes value `r i` on `v i` for all `i` in `s`. -/ +polynomial of degree `< #s` that takes value `r i` on `v i` for all `i` in `s`. -/ @[simps] def interpolate (s : Finset ι) (v : ι → F) : (ι → F) →ₗ[F] F[X] where toFun r := ∑ i ∈ s, C (r i) * Lagrange.basis s v i @@ -316,7 +317,7 @@ theorem eval_interpolate_at_node (hvs : Set.InjOn v s) (hi : i ∈ s) : rw [eval_basis_of_ne (mem_erase.mp H).1 hi, mul_zero] theorem degree_interpolate_le (hvs : Set.InjOn v s) : - (interpolate s v r).degree ≤ ↑(s.card - 1) := by + (interpolate s v r).degree ≤ ↑(#s - 1) := by refine (degree_sum_le _ _).trans ?_ rw [Finset.sup_le_iff] intro i hi @@ -326,7 +327,7 @@ theorem degree_interpolate_le (hvs : Set.InjOn v s) : · rw [degree_C hr, zero_add] -- Porting note: Added `Nat.cast_withBot` rewrites -theorem degree_interpolate_lt (hvs : Set.InjOn v s) : (interpolate s v r).degree < s.card := by +theorem degree_interpolate_lt (hvs : Set.InjOn v s) : (interpolate s v r).degree < #s := by rw [Nat.cast_withBot] rcases eq_empty_or_nonempty s with (rfl | h) · rw [interpolate_empty, degree_zero, card_empty] @@ -336,7 +337,7 @@ theorem degree_interpolate_lt (hvs : Set.InjOn v s) : (interpolate s v r).degree exact Nat.sub_lt (Nonempty.card_pos h) zero_lt_one theorem degree_interpolate_erase_lt (hvs : Set.InjOn v s) (hi : i ∈ s) : - (interpolate (s.erase i) v r).degree < ↑(s.card - 1) := by + (interpolate (s.erase i) v r).degree < ↑(#s - 1) := by rw [← Finset.card_erase_of_mem hi] exact degree_interpolate_lt _ (Set.InjOn.mono (coe_subset.mpr (erase_subset _ _)) hvs) @@ -352,12 +353,12 @@ theorem interpolate_eq_iff_values_eq_on (hvs : Set.InjOn v s) : interpolate s v r = interpolate s v r' ↔ ∀ i ∈ s, r i = r' i := ⟨values_eq_on_of_interpolate_eq _ _ hvs, interpolate_eq_of_values_eq_on _ _⟩ -theorem eq_interpolate {f : F[X]} (hvs : Set.InjOn v s) (degree_f_lt : f.degree < s.card) : +theorem eq_interpolate {f : F[X]} (hvs : Set.InjOn v s) (degree_f_lt : f.degree < #s) : f = interpolate s v fun i => f.eval (v i) := eq_of_degrees_lt_of_eval_index_eq _ hvs degree_f_lt (degree_interpolate_lt _ hvs) fun _ hi => (eval_interpolate_at_node (fun x ↦ eval (v x) f) hvs hi).symm -theorem eq_interpolate_of_eval_eq {f : F[X]} (hvs : Set.InjOn v s) (degree_f_lt : f.degree < s.card) +theorem eq_interpolate_of_eval_eq {f : F[X]} (hvs : Set.InjOn v s) (degree_f_lt : f.degree < #s) (eval_f : ∀ i ∈ s, f.eval (v i) = r i) : f = interpolate s v r := by rw [eq_interpolate hvs degree_f_lt] exact interpolate_eq_of_values_eq_on _ _ eval_f @@ -366,7 +367,7 @@ theorem eq_interpolate_of_eval_eq {f : F[X]} (hvs : Set.InjOn v s) (degree_f_lt unique polynomial of `degree < Fintype.card ι` which takes the value of the `r i` on the `v i`. -/ theorem eq_interpolate_iff {f : F[X]} (hvs : Set.InjOn v s) : - (f.degree < s.card ∧ ∀ i ∈ s, eval (v i) f = r i) ↔ f = interpolate s v r := by + (f.degree < #s ∧ ∀ i ∈ s, eval (v i) f = r i) ↔ f = interpolate s v r := by constructor <;> intro h · exact eq_interpolate_of_eval_eq _ hvs h.1 h.2 · rw [h] @@ -374,7 +375,7 @@ theorem eq_interpolate_iff {f : F[X]} (hvs : Set.InjOn v s) : /-- Lagrange interpolation induces isomorphism between functions from `s` and polynomials of degree less than `Fintype.card ι`. -/ -def funEquivDegreeLT (hvs : Set.InjOn v s) : degreeLT F s.card ≃ₗ[F] s → F where +def funEquivDegreeLT (hvs : Set.InjOn v s) : degreeLT F #s ≃ₗ[F] s → F where toFun f i := f.1.eval (v i) map_add' _ _ := funext fun _ => eval_add map_smul' c f := funext <| by simp @@ -399,15 +400,15 @@ theorem interpolate_eq_sum_interpolate_insert_sdiff (hvt : Set.InjOn v t) (hs : interpolate t v r = ∑ i ∈ s, interpolate (insert i (t \ s)) v r * Lagrange.basis s v i := by symm refine eq_interpolate_of_eval_eq _ hvt (lt_of_le_of_lt (degree_sum_le _ _) ?_) fun i hi => ?_ - · simp_rw [Nat.cast_withBot, Finset.sup_lt_iff (WithBot.bot_lt_coe t.card), degree_mul] + · simp_rw [Nat.cast_withBot, Finset.sup_lt_iff (WithBot.bot_lt_coe #t), degree_mul] intro i hi - have hs : 1 ≤ s.card := Nonempty.card_pos ⟨_, hi⟩ - have hst' : s.card ≤ t.card := card_le_card hst - have H : t.card = 1 + (t.card - s.card) + (s.card - 1) := by + have hs : 1 ≤ #s := Nonempty.card_pos ⟨_, hi⟩ + have hst' : #s ≤ #t := card_le_card hst + have H : #t = 1 + (#t - #s) + (#s - 1) := by rw [add_assoc, tsub_add_tsub_cancel hst' hs, ← add_tsub_assoc_of_le (hs.trans hst'), Nat.succ_add_sub_one, zero_add] rw [degree_basis (Set.InjOn.mono hst hvt) hi, H, WithBot.coe_add, Nat.cast_withBot, - WithBot.add_lt_add_iff_right (@WithBot.coe_ne_bot _ (s.card - 1))] + WithBot.add_lt_add_iff_right (@WithBot.coe_ne_bot _ (#s - 1))] convert degree_interpolate_lt _ (hvt.mono (coe_subset.mpr (insert_subset_iff.mpr ⟨hst hi, sdiff_subset⟩))) rw [card_insert_of_not_mem (not_mem_sdiff_of_mem_right hi), card_sdiff hst, add_comm] @@ -471,7 +472,7 @@ theorem nodal_empty : nodal ∅ v = 1 := by rfl @[simp] -theorem natDegree_nodal [Nontrivial R] : (nodal s v).natDegree = s.card := by +theorem natDegree_nodal [Nontrivial R] : (nodal s v).natDegree = #s := by simp_rw [nodal, natDegree_prod_of_monic (h := fun i _ => monic_X_sub_C (v i)), natDegree_X_sub_C, sum_const, smul_eq_mul, mul_one] @@ -482,7 +483,7 @@ rcases s.eq_empty_or_nonempty with (rfl | h) simp only [natDegree_nodal, h.card_pos] @[simp] -theorem degree_nodal [Nontrivial R] : (nodal s v).degree = s.card := by +theorem degree_nodal [Nontrivial R] : (nodal s v).degree = #s := by simp_rw [degree_eq_natDegree nodal_ne_zero, natDegree_nodal] theorem nodal_monic : (nodal s v).Monic := diff --git a/Mathlib/LinearAlgebra/Matrix/AbsoluteValue.lean b/Mathlib/LinearAlgebra/Matrix/AbsoluteValue.lean index 1af383318a9ae..201767b2c0744 100644 --- a/Mathlib/LinearAlgebra/Matrix/AbsoluteValue.lean +++ b/Mathlib/LinearAlgebra/Matrix/AbsoluteValue.lean @@ -50,19 +50,19 @@ theorem det_le {A : Matrix n n R} {abv : AbsoluteValue R S} {x : S} (hx : ∀ i theorem det_sum_le {ι : Type*} (s : Finset ι) {A : ι → Matrix n n R} {abv : AbsoluteValue R S} {x : S} (hx : ∀ k i j, abv (A k i j) ≤ x) : abv (det (∑ k ∈ s, A k)) ≤ - Nat.factorial (Fintype.card n) • (Finset.card s • x) ^ Fintype.card n := + Nat.factorial (Fintype.card n) • (#s• x) ^ Fintype.card n := det_le fun i j => calc abv ((∑ k ∈ s, A k) i j) = abv (∑ k ∈ s, A k i j) := by simp only [sum_apply] _ ≤ ∑ k ∈ s, abv (A k i j) := abv.sum_le _ _ _ ≤ ∑ _k ∈ s, x := sum_le_sum fun k _ => hx k i j - _ = s.card • x := sum_const _ + _ = #s • x := sum_const _ theorem det_sum_smul_le {ι : Type*} (s : Finset ι) {c : ι → R} {A : ι → Matrix n n R} {abv : AbsoluteValue R S} {x : S} (hx : ∀ k i j, abv (A k i j) ≤ x) {y : S} (hy : ∀ k, abv (c k) ≤ y) : abv (det (∑ k ∈ s, c k • A k)) ≤ - Nat.factorial (Fintype.card n) • (Finset.card s • y * x) ^ Fintype.card n := by + Nat.factorial (Fintype.card n) • (#s • y * x) ^ Fintype.card n := by simpa only [smul_mul_assoc] using det_sum_le s fun k i j => calc diff --git a/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean b/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean index c33327a8858ae..2780bc5cd41ea 100644 --- a/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean +++ b/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean @@ -131,12 +131,9 @@ theorem det_mul (M N : Matrix n n R) : det (M * N) = det M * det N := det (M * N) = ∑ p : n → n, ∑ σ : Perm n, ε σ * ∏ i, M (σ i) (p i) * N (p i) i := by simp only [det_apply', mul_apply, prod_univ_sum, mul_sum, Fintype.piFinset_univ] rw [Finset.sum_comm] - _ = - ∑ p ∈ (@univ (n → n) _).filter Bijective, - ∑ σ : Perm n, ε σ * ∏ i, M (σ i) (p i) * N (p i) i := - (Eq.symm <| - sum_subset (filter_subset _ _) fun f _ hbij => - det_mul_aux <| by simpa only [true_and, mem_filter, mem_univ] using hbij) + _ = ∑ p : n → n with Bijective p, ∑ σ : Perm n, ε σ * ∏ i, M (σ i) (p i) * N (p i) i := by + refine (sum_subset (filter_subset _ _) fun f _ hbij ↦ det_mul_aux ?_).symm + simpa only [true_and, mem_filter, mem_univ] using hbij _ = ∑ τ : Perm n, ∑ σ : Perm n, ε σ * ∏ i, M (σ i) (τ i) * N (τ i) i := sum_bij (fun p h ↦ Equiv.ofBijective p (mem_filter.1 h).2) (fun _ _ ↦ mem_univ _) (fun _ _ _ _ h ↦ by injection h) @@ -543,8 +540,7 @@ theorem det_blockDiagonal {o : Type*} [Fintype o] [DecidableEq o] (M : o → Mat simp_rw [Finset.prod_attach_univ, Finset.univ_pi_univ] -- We claim that the only permutations contributing to the sum are those that -- preserve their second component. - let preserving_snd : Finset (Equiv.Perm (n × o)) := - Finset.univ.filter fun σ => ∀ x, (σ x).snd = x.snd + let preserving_snd : Finset (Equiv.Perm (n × o)) := {σ | ∀ x, (σ x).snd = x.snd} have mem_preserving_snd : ∀ {σ : Equiv.Perm (n × o)}, σ ∈ preserving_snd ↔ ∀ x, (σ x).snd = x.snd := fun {σ} => Finset.mem_filter.trans ⟨fun h => h.2, fun h => ⟨Finset.mem_univ _, h⟩⟩ diff --git a/Mathlib/LinearAlgebra/Multilinear/Basic.lean b/Mathlib/LinearAlgebra/Multilinear/Basic.lean index f37192686c399..39f1d7e414a2b 100644 --- a/Mathlib/LinearAlgebra/Multilinear/Basic.lean +++ b/Mathlib/LinearAlgebra/Multilinear/Basic.lean @@ -71,8 +71,7 @@ something similar, but of the form `Fin.decidableEq n = _inst`, which is much ea since `_inst` is a free variable and so the equality can just be substituted. -/ - -open Function Fin Set +open Fin Function Finset Set universe uR uS uι v v' v₁ v₂ v₃ @@ -280,7 +279,7 @@ the other ones equal to a given value `z`. It is denoted by `f.restr s hk z`, wh proof that the cardinality of `s` is `k`. The implicit identification between `Fin k` and `s` that we use is the canonical (increasing) bijection. -/ def restr {k n : ℕ} (f : MultilinearMap R (fun _ : Fin n => M') M₂) (s : Finset (Fin n)) - (hk : s.card = k) (z : M') : MultilinearMap R (fun _ : Fin k => M') M₂ where + (hk : #s = k) (z : M') : MultilinearMap R (fun _ : Fin k => M') M₂ where toFun v := f fun j => if h : j ∈ s then v ((DFunLike.coe (s.orderIsoOfFin hk).symm) ⟨j, h⟩) else z /- Porting note: The proofs of the following two lemmas used to only use `erw` followed by `simp`, but it seems `erw` no longer unfolds or unifies well enough to work without more help. -/ @@ -445,7 +444,7 @@ open Fintype Finset `r n ∈ Aₙ`. This follows from multilinearity by expanding successively with respect to each coordinate. Here, we give an auxiliary statement tailored for an inductive proof. Use instead `map_sum_finset`. -/ -theorem map_sum_finset_aux [DecidableEq ι] [Fintype ι] {n : ℕ} (h : (∑ i, (A i).card) = n) : +theorem map_sum_finset_aux [DecidableEq ι] [Fintype ι] {n : ℕ} (h : (∑ i, #(A i)) = n) : (f fun i => ∑ j ∈ A i, g i j) = ∑ r ∈ piFinset A, f fun i => g i (r i) := by letI := fun i => Classical.decEq (α i) induction' n using Nat.strong_induction_on with n IH generalizing A @@ -457,11 +456,11 @@ theorem map_sum_finset_aux [DecidableEq ι] [Fintype ι] {n : ℕ} (h : (∑ i, push_neg at Ai_empty -- Otherwise, if all sets are at most singletons, then they are exactly singletons and the result -- is again straightforward - by_cases Ai_singleton : ∀ i, (A i).card ≤ 1 - · have Ai_card : ∀ i, (A i).card = 1 := by + by_cases Ai_singleton : ∀ i, #(A i) ≤ 1 + · have Ai_card : ∀ i, #(A i) = 1 := by intro i - have pos : Finset.card (A i) ≠ 0 := by simp [Finset.card_eq_zero, Ai_empty i] - have : Finset.card (A i) ≤ 1 := Ai_singleton i + have pos : #(A i) ≠ 0 := by simp [Finset.card_eq_zero, Ai_empty i] + have : #(A i) ≤ 1 := Ai_singleton i exact le_antisymm this (Nat.succ_le_of_lt (_root_.pos_iff_ne_zero.mpr pos)) have : ∀ r : ∀ i, α i, r ∈ piFinset A → (f fun i => g i (r i)) = f fun i => ∑ j ∈ A i, g i j := by @@ -480,7 +479,7 @@ theorem map_sum_finset_aux [DecidableEq ι] [Fintype ι] {n : ℕ} (h : (∑ i, -- for `i ≠ i₀`, apply the inductive assumption to `B` and `C`, and add up the corresponding -- parts to get the sum for `A`. push_neg at Ai_singleton - obtain ⟨i₀, hi₀⟩ : ∃ i, 1 < (A i).card := Ai_singleton + obtain ⟨i₀, hi₀⟩ : ∃ i, 1 < #(A i) := Ai_singleton obtain ⟨j₁, j₂, _, hj₂, _⟩ : ∃ j₁ j₂, j₁ ∈ A i₀ ∧ j₂ ∈ A i₀ ∧ j₁ ≠ j₂ := Finset.one_lt_card_iff.1 hi₀ let B := Function.update A i₀ (A i₀ \ {j₂}) @@ -535,10 +534,8 @@ theorem map_sum_finset_aux [DecidableEq ι] [Fintype ι] {n : ℕ} (h : (∑ i, · simp only [C, hi, update_noteq, Ne, not_false_iff] -- Express the inductive assumption for `B` have Brec : (f fun i => ∑ j ∈ B i, g i j) = ∑ r ∈ piFinset B, f fun i => g i (r i) := by - have : (∑ i, Finset.card (B i)) < ∑ i, Finset.card (A i) := by - refine - Finset.sum_lt_sum (fun i _ => Finset.card_le_card (B_subset_A i)) - ⟨i₀, Finset.mem_univ _, ?_⟩ + have : ∑ i, #(B i) < ∑ i, #(A i) := by + refine sum_lt_sum (fun i _ => card_le_card (B_subset_A i)) ⟨i₀, mem_univ _, ?_⟩ have : {j₂} ⊆ A i₀ := by simp [hj₂] simp only [B, Finset.card_sdiff this, Function.update_same, Finset.card_singleton] exact Nat.pred_lt (ne_of_gt (lt_trans Nat.zero_lt_one hi₀)) @@ -546,7 +543,7 @@ theorem map_sum_finset_aux [DecidableEq ι] [Fintype ι] {n : ℕ} (h : (∑ i, exact IH _ this B rfl -- Express the inductive assumption for `C` have Crec : (f fun i => ∑ j ∈ C i, g i j) = ∑ r ∈ piFinset C, f fun i => g i (r i) := by - have : (∑ i, Finset.card (C i)) < ∑ i, Finset.card (A i) := + have : (∑ i, #(C i)) < ∑ i, #(A i) := Finset.sum_lt_sum (fun i _ => Finset.card_le_card (C_subset_A i)) ⟨i₀, Finset.mem_univ _, by simp [C, hi₀]⟩ rw [h] at this @@ -1319,10 +1316,9 @@ lemma map_piecewise_sub_map_piecewise [LinearOrder ι] (a b v : (i : ι) → M open Finset in lemma map_add_eq_map_add_linearDeriv_add [DecidableEq ι] [Fintype ι] (x h : (i : ι) → M₁ i) : - f (x + h) = f x + f.linearDeriv x h + - ∑ s ∈ univ.powerset.filter (2 ≤ ·.card), f (s.piecewise h x) := by + f (x + h) = f x + f.linearDeriv x h + ∑ s with 2 ≤ #s, f (s.piecewise h x) := by rw [add_comm, map_add_univ, ← Finset.powerset_univ, - ← sum_filter_add_sum_filter_not _ (2 ≤ ·.card)] + ← sum_filter_add_sum_filter_not _ (2 ≤ #·)] simp_rw [not_le, Nat.lt_succ, le_iff_lt_or_eq (b := 1), Nat.lt_one_iff, filter_or, ← powersetCard_eq_filter, sum_union (univ.pairwise_disjoint_powersetCard zero_ne_one), powersetCard_zero, powersetCard_one, sum_singleton, Finset.piecewise_empty, sum_map, @@ -1334,7 +1330,7 @@ at two points "close to `x`" in terms of the "derivative" of the multilinear map and of "second-order" terms. -/ lemma map_add_sub_map_add_sub_linearDeriv [DecidableEq ι] [Fintype ι] (x h h' : (i : ι) → M₁ i) : f (x + h) - f (x + h') - f.linearDeriv x (h - h') = - ∑ s ∈ univ.powerset.filter (2 ≤ ·.card), (f (s.piecewise h x) - f (s.piecewise h' x)) := by + ∑ s with 2 ≤ #s, (f (s.piecewise h x) - f (s.piecewise h' x)) := by simp_rw [map_add_eq_map_add_linearDeriv_add, add_assoc, add_sub_add_comm, sub_self, zero_add, ← LinearMap.map_sub, add_sub_cancel_left, sum_sub_distrib] @@ -1679,7 +1675,7 @@ variable (R M₂ M') `l`, then the space of multilinear maps on `fun i : Fin n => M'` is isomorphic to the space of multilinear maps on `fun i : Fin k => M'` taking values in the space of multilinear maps on `fun i : Fin l => M'`. -/ -def curryFinFinset {k l n : ℕ} {s : Finset (Fin n)} (hk : s.card = k) (hl : sᶜ.card = l) : +def curryFinFinset {k l n : ℕ} {s : Finset (Fin n)} (hk : #s = k) (hl : #sᶜ = l) : MultilinearMap R (fun _ : Fin n => M') M₂ ≃ₗ[R] MultilinearMap R (fun _ : Fin k => M') (MultilinearMap R (fun _ : Fin l => M') M₂) := (domDomCongrLinearEquiv R R M' M₂ (finSumEquivOfFinset hk hl).symm).trans @@ -1688,15 +1684,15 @@ def curryFinFinset {k l n : ℕ} {s : Finset (Fin n)} (hk : s.card = k) (hl : s variable {R M₂ M'} @[simp] -theorem curryFinFinset_apply {k l n : ℕ} {s : Finset (Fin n)} (hk : s.card = k) (hl : sᶜ.card = l) +theorem curryFinFinset_apply {k l n : ℕ} {s : Finset (Fin n)} (hk : #s = k) (hl : #sᶜ = l) (f : MultilinearMap R (fun _ : Fin n => M') M₂) (mk : Fin k → M') (ml : Fin l → M') : curryFinFinset R M₂ M' hk hl f mk ml = f fun i => Sum.elim mk ml ((finSumEquivOfFinset hk hl).symm i) := rfl @[simp] -theorem curryFinFinset_symm_apply {k l n : ℕ} {s : Finset (Fin n)} (hk : s.card = k) - (hl : sᶜ.card = l) +theorem curryFinFinset_symm_apply {k l n : ℕ} {s : Finset (Fin n)} (hk : #s = k) + (hl : #sᶜ = l) (f : MultilinearMap R (fun _ : Fin k => M') (MultilinearMap R (fun _ : Fin l => M') M₂)) (m : Fin n → M') : (curryFinFinset R M₂ M' hk hl).symm f m = @@ -1705,8 +1701,8 @@ theorem curryFinFinset_symm_apply {k l n : ℕ} {s : Finset (Fin n)} (hk : s.car rfl -- @[simp] -- Porting note: simpNF linter, lhs simplifies, added aux version below -theorem curryFinFinset_symm_apply_piecewise_const {k l n : ℕ} {s : Finset (Fin n)} (hk : s.card = k) - (hl : sᶜ.card = l) +theorem curryFinFinset_symm_apply_piecewise_const {k l n : ℕ} {s : Finset (Fin n)} (hk : #s = k) + (hl : #sᶜ = l) (f : MultilinearMap R (fun _ : Fin k => M') (MultilinearMap R (fun _ : Fin l => M') M₂)) (x y : M') : (curryFinFinset R M₂ M' hk hl).symm f (s.piecewise (fun _ => x) fun _ => y) = @@ -1721,7 +1717,7 @@ theorem curryFinFinset_symm_apply_piecewise_const {k l n : ℕ} {s : Finset (Fin @[simp] theorem curryFinFinset_symm_apply_piecewise_const_aux {k l n : ℕ} {s : Finset (Fin n)} - (hk : s.card = k) (hl : sᶜ.card = l) + (hk : #s = k) (hl : #sᶜ = l) (f : MultilinearMap R (fun _ : Fin k => M') (MultilinearMap R (fun _ : Fin l => M') M₂)) (x y : M') : ((⇑f fun _ => x) (fun i => (Finset.piecewise s (fun _ => x) (fun _ => y) @@ -1732,15 +1728,15 @@ theorem curryFinFinset_symm_apply_piecewise_const_aux {k l n : ℕ} {s : Finset exact this @[simp] -theorem curryFinFinset_symm_apply_const {k l n : ℕ} {s : Finset (Fin n)} (hk : s.card = k) - (hl : sᶜ.card = l) +theorem curryFinFinset_symm_apply_const {k l n : ℕ} {s : Finset (Fin n)} (hk : #s = k) + (hl : #sᶜ = l) (f : MultilinearMap R (fun _ : Fin k => M') (MultilinearMap R (fun _ : Fin l => M') M₂)) (x : M') : ((curryFinFinset R M₂ M' hk hl).symm f fun _ => x) = f (fun _ => x) fun _ => x := rfl -- @[simp] -- Porting note: simpNF, lhs simplifies, added aux version below -theorem curryFinFinset_apply_const {k l n : ℕ} {s : Finset (Fin n)} (hk : s.card = k) - (hl : sᶜ.card = l) (f : MultilinearMap R (fun _ : Fin n => M') M₂) (x y : M') : +theorem curryFinFinset_apply_const {k l n : ℕ} {s : Finset (Fin n)} (hk : #s = k) + (hl : #sᶜ = l) (f : MultilinearMap R (fun _ : Fin n => M') M₂) (x y : M') : (curryFinFinset R M₂ M' hk hl f (fun _ => x) fun _ => y) = f (s.piecewise (fun _ => x) fun _ => y) := by refine (curryFinFinset_symm_apply_piecewise_const hk hl _ _ _).symm.trans ?_ @@ -1748,8 +1744,8 @@ theorem curryFinFinset_apply_const {k l n : ℕ} {s : Finset (Fin n)} (hk : s.ca rw [LinearEquiv.symm_apply_apply] @[simp] -theorem curryFinFinset_apply_const_aux {k l n : ℕ} {s : Finset (Fin n)} (hk : s.card = k) - (hl : sᶜ.card = l) (f : MultilinearMap R (fun _ : Fin n => M') M₂) (x y : M') : +theorem curryFinFinset_apply_const_aux {k l n : ℕ} {s : Finset (Fin n)} (hk : #s = k) + (hl : #sᶜ = l) (f : MultilinearMap R (fun _ : Fin n => M') M₂) (x y : M') : (f fun i => Sum.elim (fun _ => x) (fun _ => y) ((⇑ (Equiv.symm (finSumEquivOfFinset hk hl))) i)) = f (s.piecewise (fun _ => x) fun _ => y) := by rw [← curryFinFinset_apply] diff --git a/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean b/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean index 39bc41327d511..99076960053fe 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean @@ -329,7 +329,7 @@ theorem choose_exists_companion : Q.exists_companion.choose = polarBilin Q := protected theorem map_sum {ι} [DecidableEq ι] (Q : QuadraticMap R M N) (s : Finset ι) (f : ι → M) : Q (∑ i ∈ s, f i) = ∑ i ∈ s, Q (f i) + - ∑ ij in s.sym2.filter (¬ ·.IsDiag), + ∑ ij ∈ s.sym2 with ¬ ij.IsDiag, Sym2.lift ⟨fun i j => polar Q (f i) (f j), fun _ _ => polar_comm _ _ _⟩ ij := by induction s using Finset.cons_induction with | empty => simp diff --git a/Mathlib/Logic/Denumerable.lean b/Mathlib/Logic/Denumerable.lean index cf0e64c97bd1d..f96dccd8f839c 100644 --- a/Mathlib/Logic/Denumerable.lean +++ b/Mathlib/Logic/Denumerable.lean @@ -30,7 +30,7 @@ class Denumerable (α : Type*) extends Encodable α where /-- `decode` and `encode` are inverses. -/ decode_inv : ∀ n, ∃ a ∈ decode n, encode a = n -open Nat +open Finset Nat namespace Denumerable @@ -274,7 +274,7 @@ private def toFunAux (x : s) : ℕ := (List.range x).countP (· ∈ s) private theorem toFunAux_eq {s : Set ℕ} [DecidablePred (· ∈ s)] (x : s) : - toFunAux x = ((Finset.range x).filter (· ∈ s)).card := by + toFunAux x = #{y ∈ Finset.range x | y ∈ s} := by rw [toFunAux, List.countP_eq_length_filter] rfl @@ -288,9 +288,9 @@ private theorem right_inverse_aux : ∀ n, toFunAux (ofNat s n) = n exact bot_le.not_lt (show (⟨n, hn.2⟩ : s) < ⊥ from hn.1) | n + 1 => by have ih : toFunAux (ofNat s n) = n := right_inverse_aux n - have h₁ : (ofNat s n : ℕ) ∉ (range (ofNat s n)).filter (· ∈ s) := by simp - have h₂ : (range (succ (ofNat s n))).filter (· ∈ s) = - insert ↑(ofNat s n) ((range (ofNat s n)).filter (· ∈ s)) := by + have h₁ : (ofNat s n : ℕ) ∉ {x ∈ range (ofNat s n) | x ∈ s} := by simp + have h₂ : {x ∈ range (succ (ofNat s n)) | x ∈ s} = + insert ↑(ofNat s n) {x ∈ range (ofNat s n) | x ∈ s} := by simp only [Finset.ext_iff, mem_insert, mem_range, mem_filter] exact fun m => ⟨fun h => by diff --git a/Mathlib/MeasureTheory/Function/SimpleFunc.lean b/Mathlib/MeasureTheory/Function/SimpleFunc.lean index 6a1c98e568221..0a08303ca10a2 100644 --- a/Mathlib/MeasureTheory/Function/SimpleFunc.lean +++ b/Mathlib/MeasureTheory/Function/SimpleFunc.lean @@ -269,13 +269,13 @@ theorem map_const (g : β → γ) (b : β) : (const α b).map g = const α (g b) rfl theorem map_preimage (f : α →ₛ β) (g : β → γ) (s : Set γ) : - f.map g ⁻¹' s = f ⁻¹' ↑(f.range.filter fun b => g b ∈ s) := by + f.map g ⁻¹' s = f ⁻¹' ↑{b ∈ f.range | g b ∈ s} := by simp only [coe_range, sep_mem_eq, coe_map, Finset.coe_filter, ← mem_preimage, inter_comm, preimage_inter_range, ← Finset.mem_coe] exact preimage_comp theorem map_preimage_singleton (f : α →ₛ β) (g : β → γ) (c : γ) : - f.map g ⁻¹' {c} = f ⁻¹' ↑(f.range.filter fun b => g b = c) := + f.map g ⁻¹' {c} = f ⁻¹' ↑{b ∈ f.range | g b = c} := map_preimage _ _ _ /-- Composition of a `SimpleFun` and a measurable function is a `SimpleFunc`. -/ @@ -985,7 +985,7 @@ section FinMeasSupp open Finset Function theorem support_eq [MeasurableSpace α] [Zero β] (f : α →ₛ β) : - support f = ⋃ y ∈ f.range.filter fun y => y ≠ 0, f ⁻¹' {y} := + support f = ⋃ y ∈ {y ∈ f.range | y ≠ 0}, f ⁻¹' {y} := Set.ext fun x => by simp only [mem_support, Set.mem_preimage, mem_filter, mem_range_self, true_and, exists_prop, mem_iUnion, Set.mem_range, mem_singleton_iff, exists_eq_right'] diff --git a/Mathlib/MeasureTheory/Integral/Bochner.lean b/Mathlib/MeasureTheory/Integral/Bochner.lean index 3c5f56fce5d8b..ada8eefa261f7 100644 --- a/Mathlib/MeasureTheory/Integral/Bochner.lean +++ b/Mathlib/MeasureTheory/Integral/Bochner.lean @@ -295,12 +295,12 @@ theorem integral_eq {m : MeasurableSpace α} (μ : Measure α) (f : α →ₛ F) theorem integral_eq_sum_filter [DecidablePred fun x : F => x ≠ 0] {m : MeasurableSpace α} (f : α →ₛ F) (μ : Measure α) : - f.integral μ = ∑ x ∈ f.range.filter fun x => x ≠ 0, (μ (f ⁻¹' {x})).toReal • x := by + f.integral μ = ∑ x ∈ {x ∈ f.range | x ≠ 0}, (μ (f ⁻¹' {x})).toReal • x := by simp_rw [integral_def, setToSimpleFunc_eq_sum_filter, weightedSMul_apply] /-- The Bochner integral is equal to a sum over any set that includes `f.range` (except `0`). -/ theorem integral_eq_sum_of_subset [DecidablePred fun x : F => x ≠ 0] {f : α →ₛ F} {s : Finset F} - (hs : (f.range.filter fun x => x ≠ 0) ⊆ s) : + (hs : {x ∈ f.range | x ≠ 0} ⊆ s) : f.integral μ = ∑ x ∈ s, (μ (f ⁻¹' {x})).toReal • x := by rw [SimpleFunc.integral_eq_sum_filter, Finset.sum_subset hs] rintro x - hx; rw [Finset.mem_filter, not_and_or, Ne, Classical.not_not] at hx diff --git a/Mathlib/MeasureTheory/Integral/SetToL1.lean b/Mathlib/MeasureTheory/Integral/SetToL1.lean index 3d16c95fed2c1..456069af75266 100644 --- a/Mathlib/MeasureTheory/Integral/SetToL1.lean +++ b/Mathlib/MeasureTheory/Integral/SetToL1.lean @@ -277,9 +277,8 @@ theorem setToSimpleFunc_zero_apply {m : MeasurableSpace α} (T : Set α → F cases isEmpty_or_nonempty α <;> simp [setToSimpleFunc] theorem setToSimpleFunc_eq_sum_filter [DecidablePred fun x ↦ x ≠ (0 : F)] - {m : MeasurableSpace α} (T : Set α → F →L[ℝ] F') - (f : α →ₛ F) : - setToSimpleFunc T f = ∑ x ∈ f.range.filter fun x => x ≠ 0, (T (f ⁻¹' {x})) x := by + {m : MeasurableSpace α} (T : Set α → F →L[ℝ] F') (f : α →ₛ F) : + setToSimpleFunc T f = ∑ x ∈ f.range with x ≠ 0, T (f ⁻¹' {x}) x := by symm refine sum_filter_of_ne fun x _ => mt fun hx0 => ?_ rw [hx0] @@ -301,13 +300,13 @@ theorem map_setToSimpleFunc (T : Set α → F →L[ℝ] F') (h_add : FinMeasAddi rw [mem_filter] at hx rw [hx.2, ContinuousLinearMap.map_zero] have h_left_eq : - T (map g f ⁻¹' {g (f a)}) (g (f a)) = - T (f ⁻¹' (f.range.filter fun b => g b = g (f a))) (g (f a)) := by + T (map g f ⁻¹' {g (f a)}) (g (f a)) + = T (f ⁻¹' ({b ∈ f.range | g b = g (f a)} : Finset _)) (g (f a)) := by congr; rw [map_preimage_singleton] rw [h_left_eq] have h_left_eq' : - T (f ⁻¹' (filter (fun b : G => g b = g (f a)) f.range)) (g (f a)) = - T (⋃ y ∈ filter (fun b : G => g b = g (f a)) f.range, f ⁻¹' {y}) (g (f a)) := by + T (f ⁻¹' ({b ∈ f.range | g b = g (f a)} : Finset _)) (g (f a)) + = T (⋃ y ∈ {b ∈ f.range | g b = g (f a)}, f ⁻¹' {y}) (g (f a)) := by congr; rw [← Finset.set_biUnion_preimage_singleton] rw [h_left_eq'] rw [h_add.map_iUnion_fin_meas_set_eq_sum T T_empty] @@ -381,7 +380,7 @@ theorem setToSimpleFunc_add_left' (T T' T'' : Set α → E →L[ℝ] F) classical simp_rw [setToSimpleFunc_eq_sum_filter] suffices - ∀ x ∈ filter (fun x : E => x ≠ 0) f.range, T'' (f ⁻¹' {x}) = T (f ⁻¹' {x}) + T' (f ⁻¹' {x}) by + ∀ x ∈ {x ∈ f.range | x ≠ 0}, T'' (f ⁻¹' {x}) = T (f ⁻¹' {x}) + T' (f ⁻¹' {x}) by rw [← sum_add_distrib] refine Finset.sum_congr rfl fun x hx => ?_ rw [this x hx] @@ -402,7 +401,7 @@ theorem setToSimpleFunc_smul_left' (T T' : Set α → E →L[ℝ] F') (c : ℝ) setToSimpleFunc T' f = c • setToSimpleFunc T f := by classical simp_rw [setToSimpleFunc_eq_sum_filter] - suffices ∀ x ∈ filter (fun x : E => x ≠ 0) f.range, T' (f ⁻¹' {x}) = c • T (f ⁻¹' {x}) by + suffices ∀ x ∈ {x ∈ f.range | x ≠ 0}, T' (f ⁻¹' {x}) = c • T (f ⁻¹' {x}) by rw [smul_sum] refine Finset.sum_congr rfl fun x hx => ?_ rw [this x hx] diff --git a/Mathlib/NumberTheory/ArithmeticFunction.lean b/Mathlib/NumberTheory/ArithmeticFunction.lean index 71c09ddd1baad..f954d1b9acb24 100644 --- a/Mathlib/NumberTheory/ArithmeticFunction.lean +++ b/Mathlib/NumberTheory/ArithmeticFunction.lean @@ -822,7 +822,7 @@ theorem sigma_one_apply_prime_pow {p i : ℕ} (hp : p.Prime) : σ 1 (p ^ i) = ∑ k in .range (i + 1), p ^ k := by simp [sigma_apply_prime_pow hp] -theorem sigma_zero_apply (n : ℕ) : σ 0 n = (divisors n).card := by simp [sigma_apply] +theorem sigma_zero_apply (n : ℕ) : σ 0 n = #n.divisors := by simp [sigma_apply] theorem sigma_zero_apply_prime_pow {p i : ℕ} (hp : p.Prime) : σ 0 (p ^ i) = i + 1 := by simp [sigma_apply_prime_pow hp] @@ -1286,13 +1286,13 @@ theorem prod_eq_iff_prod_pow_moebius_eq_on_of_nonzero [CommGroupWithZero R] end SpecialFunctions theorem _root_.Nat.card_divisors {n : ℕ} (hn : n ≠ 0) : - n.divisors.card = n.primeFactors.prod (n.factorization · + 1) := by + #n.divisors = n.primeFactors.prod (n.factorization · + 1) := by rw [← sigma_zero_apply, isMultiplicative_sigma.multiplicative_factorization _ hn] exact Finset.prod_congr n.support_factorization fun _ h => sigma_zero_apply_prime_pow <| Nat.prime_of_mem_primeFactors h @[deprecated (since := "2024-06-09")] theorem card_divisors (n : ℕ) (hn : n ≠ 0) : - n.divisors.card = n.primeFactors.prod (n.factorization · + 1) := Nat.card_divisors hn + #n.divisors = n.primeFactors.prod (n.factorization · + 1) := Nat.card_divisors hn theorem _root_.Nat.sum_divisors {n : ℕ} (hn : n ≠ 0) : ∑ d ∈ n.divisors, d = ∏ p ∈ n.primeFactors, ∑ k ∈ .range (n.factorization p + 1), p ^ k := by @@ -1307,7 +1307,7 @@ namespace Nat.Coprime open ArithmeticFunction theorem card_divisors_mul {m n : ℕ} (hmn : m.Coprime n) : - (m * n).divisors.card = m.divisors.card * n.divisors.card := by + #(m * n).divisors = #m.divisors * #n.divisors := by simp only [← sigma_zero_apply, isMultiplicative_sigma.map_mul_of_coprime hmn] theorem sum_divisors_mul {m n : ℕ} (hmn : m.Coprime n) : diff --git a/Mathlib/NumberTheory/DiophantineApproximation.lean b/Mathlib/NumberTheory/DiophantineApproximation.lean index 1d052c7a9f81e..604980a89620e 100644 --- a/Mathlib/NumberTheory/DiophantineApproximation.lean +++ b/Mathlib/NumberTheory/DiophantineApproximation.lean @@ -113,7 +113,7 @@ theorem exists_int_int_abs_mul_sub_le (ξ : ℝ) {n : ℕ} (n_pos : 0 < n) : simpa only [neg_add_cancel_comm_assoc] using hf' · -- Porting note(https://github.com/leanprover-community/mathlib4/issues/5127): added `not_and` simp_rw [not_exists, not_and] at H - have hD : (Ico (0 : ℤ) n).card < D.card := by rw [card_Icc, card_Ico]; exact lt_add_one n + have hD : #(Ico (0 : ℤ) n) < #D := by rw [card_Icc, card_Ico]; exact lt_add_one n have hfu' : ∀ m, f m ≤ n := fun m => lt_add_one_iff.mp (floor_lt.mpr (mod_cast hfu m)) have hwd : ∀ m : ℤ, m ∈ D → f m ∈ Ico (0 : ℤ) n := fun x hx => mem_Ico.mpr diff --git a/Mathlib/NumberTheory/Divisors.lean b/Mathlib/NumberTheory/Divisors.lean index d0f8d0cc61130..e2bfeafb8e217 100644 --- a/Mathlib/NumberTheory/Divisors.lean +++ b/Mathlib/NumberTheory/Divisors.lean @@ -36,30 +36,27 @@ namespace Nat variable (n : ℕ) /-- `divisors n` is the `Finset` of divisors of `n`. As a special case, `divisors 0 = ∅`. -/ -def divisors : Finset ℕ := - Finset.filter (fun x : ℕ => x ∣ n) (Finset.Ico 1 (n + 1)) +def divisors : Finset ℕ := {d ∈ Ico 1 (n + 1) | d ∣ n} /-- `properDivisors n` is the `Finset` of divisors of `n`, other than `n`. As a special case, `properDivisors 0 = ∅`. -/ -def properDivisors : Finset ℕ := - Finset.filter (fun x : ℕ => x ∣ n) (Finset.Ico 1 n) +def properDivisors : Finset ℕ := {d ∈ Ico 1 n | d ∣ n} /-- `divisorsAntidiagonal n` is the `Finset` of pairs `(x,y)` such that `x * y = n`. As a special case, `divisorsAntidiagonal 0 = ∅`. -/ def divisorsAntidiagonal : Finset (ℕ × ℕ) := - Finset.filter (fun x => x.fst * x.snd = n) (Ico 1 (n + 1) ×ˢ Ico 1 (n + 1)) + {x ∈ Ico 1 (n + 1) ×ˢ Ico 1 (n + 1) | x.fst * x.snd = n} variable {n} @[simp] -theorem filter_dvd_eq_divisors (h : n ≠ 0) : (Finset.range n.succ).filter (· ∣ n) = n.divisors := by +theorem filter_dvd_eq_divisors (h : n ≠ 0) : {d ∈ range n.succ | d ∣ n} = n.divisors := by ext simp only [divisors, mem_filter, mem_range, mem_Ico, and_congr_left_iff, iff_and_self] exact fun ha _ => succ_le_iff.mpr (pos_of_dvd_of_pos ha h.bot_lt) @[simp] -theorem filter_dvd_eq_properDivisors (h : n ≠ 0) : - (Finset.range n).filter (· ∣ n) = n.properDivisors := by +theorem filter_dvd_eq_properDivisors (h : n ≠ 0) : {d ∈ range n | d ∣ n} = n.properDivisors := by ext simp only [properDivisors, mem_filter, mem_range, mem_Ico, and_congr_left_iff, iff_and_self] exact fun ha _ => succ_le_iff.mpr (pos_of_dvd_of_pos ha h.bot_lt) @@ -147,7 +144,7 @@ theorem divisors_subset_properDivisors {m : ℕ} (hzero : n ≠ 0) (h : m ∣ n) (lt_of_le_of_ne (divisor_le (Nat.mem_divisors.2 ⟨h, hzero⟩)) hdiff)⟩ lemma divisors_filter_dvd_of_dvd {n m : ℕ} (hn : n ≠ 0) (hm : m ∣ n) : - (n.divisors.filter (· ∣ m)) = m.divisors := by + {d ∈ n.divisors | d ∣ m} = m.divisors := by ext k simp_rw [mem_filter, mem_divisors] exact ⟨fun ⟨_, hkm⟩ ↦ ⟨hkm, ne_zero_of_dvd_ne_zero hn hm⟩, fun ⟨hk, _⟩ ↦ ⟨⟨hk.trans hm, hn⟩, hk⟩⟩ @@ -471,7 +468,7 @@ theorem prod_divisorsAntidiagonal' {M : Type*} [CommMonoid M] (f : ℕ → ℕ /-- The factors of `n` are the prime divisors -/ theorem primeFactors_eq_to_filter_divisors_prime (n : ℕ) : - n.primeFactors = (divisors n).filter Prime := by + n.primeFactors = {p ∈ divisors n | p.Prime} := by rcases n.eq_zero_or_pos with (rfl | hn) · simp · ext q @@ -481,7 +478,7 @@ theorem primeFactors_eq_to_filter_divisors_prime (n : ℕ) : alias prime_divisors_eq_to_filter_divisors_prime := primeFactors_eq_to_filter_divisors_prime lemma primeFactors_filter_dvd_of_dvd {m n : ℕ} (hn : n ≠ 0) (hmn : m ∣ n) : - n.primeFactors.filter (· ∣ m) = m.primeFactors := by + {p ∈ n.primeFactors | p ∣ m} = m.primeFactors := by simp_rw [primeFactors_eq_to_filter_divisors_prime, filter_comm, divisors_filter_dvd_of_dvd hn hmn] diff --git a/Mathlib/NumberTheory/EulerProduct/Basic.lean b/Mathlib/NumberTheory/EulerProduct/Basic.lean index 1b6b6f551ed15..be623e1cac957 100644 --- a/Mathlib/NumberTheory/EulerProduct/Basic.lean +++ b/Mathlib/NumberTheory/EulerProduct/Basic.lean @@ -81,7 +81,7 @@ lemma summable_and_hasSum_factoredNumbers_prod_filter_prime_tsum (hsum : ∀ {p : ℕ}, p.Prime → Summable (fun n : ℕ ↦ ‖f (p ^ n)‖)) (s : Finset ℕ) : Summable (fun m : factoredNumbers s ↦ ‖f m‖) ∧ HasSum (fun m : factoredNumbers s ↦ f m) - (∏ p ∈ s.filter Nat.Prime, ∑' n : ℕ, f (p ^ n)) := by + (∏ p ∈ s with p.Prime, ∑' n : ℕ, f (p ^ n)) := by induction' s using Finset.induction with p s hp ih · rw [factoredNumbers_empty] simp only [not_mem_empty, IsEmpty.forall_iff, forall_const, filter_true_of_mem, prod_empty] @@ -93,7 +93,7 @@ lemma summable_and_hasSum_factoredNumbers_prod_filter_prime_tsum equivProdNatFactoredNumbers_apply', factoredNumbers.map_prime_pow_mul hmul hpp hp] refine Summable.of_nonneg_of_le (fun _ ↦ norm_nonneg _) (fun _ ↦ norm_mul_le ..) ?_ apply Summable.mul_of_nonneg (hsum hpp) ih.1 <;> exact fun n ↦ norm_nonneg _ - · have hp' : p ∉ s.filter Nat.Prime := mt (mem_of_mem_filter p) hp + · have hp' : p ∉ {p ∈ s | p.Prime} := mt (mem_of_mem_filter p) hp rw [prod_insert hp', ← (equivProdNatFactoredNumbers hpp hp).hasSum_iff, Function.comp_def] conv => enter [1, x] @@ -108,7 +108,7 @@ include hf₁ hmul in /-- A version of `EulerProduct.summable_and_hasSum_factoredNumbers_prod_filter_prime_tsum` in terms of the value of the series. -/ lemma prod_filter_prime_tsum_eq_tsum_factoredNumbers (hsum : Summable (‖f ·‖)) (s : Finset ℕ) : - ∏ p ∈ s.filter Nat.Prime, ∑' n : ℕ, f (p ^ n) = ∑' m : factoredNumbers s, f m := + ∏ p ∈ s with p.Prime, ∑' n : ℕ, f (p ^ n) = ∑' m : factoredNumbers s, f m := (summable_and_hasSum_factoredNumbers_prod_filter_prime_tsum hf₁ hmul (fun hp ↦ hsum.comp_injective <| Nat.pow_right_injective hp.one_lt) _).2.tsum_eq.symm @@ -175,7 +175,7 @@ theorem eulerProduct_hasProd (hsum : Summable (‖f ·‖)) (hf₀ : f 0 = 0) : intro ε hε obtain ⟨N₀, hN₀⟩ := norm_tsum_factoredNumbers_sub_tsum_lt hsum.of_norm hf₀ hε refine ⟨range N₀, fun s hs ↦ ?_⟩ - have : ∏ p ∈ s, {p | Nat.Prime p}.mulIndicator F p = ∏ p ∈ s.filter Nat.Prime, F p := + have : ∏ p ∈ s, {p | Nat.Prime p}.mulIndicator F p = ∏ p ∈ s with p.Prime, F p := prod_mulIndicator_eq_prod_filter s (fun _ ↦ F) _ id rw [this, dist_eq_norm, prod_filter_prime_tsum_eq_tsum_factoredNumbers hf₁ hmul hsum, norm_sub_rev] @@ -292,10 +292,10 @@ we show that the sum involved converges absolutely. -/ lemma summable_and_hasSum_factoredNumbers_prod_filter_prime_geometric {f : ℕ →* F} (h : ∀ {p : ℕ}, p.Prime → ‖f p‖ < 1) (s : Finset ℕ) : Summable (fun m : factoredNumbers s ↦ ‖f m‖) ∧ - HasSum (fun m : factoredNumbers s ↦ f m) (∏ p ∈ s.filter Nat.Prime, (1 - f p)⁻¹) := by + HasSum (fun m : factoredNumbers s ↦ f m) (∏ p ∈ s with p.Prime, (1 - f p)⁻¹) := by have hmul {m n} (_ : Nat.Coprime m n) := f.map_mul m n have H₁ : - ∏ p ∈ s.filter Nat.Prime, ∑' n : ℕ, f (p ^ n) = ∏ p ∈ s.filter Nat.Prime, (1 - f p)⁻¹ := by + ∏ p ∈ s with p.Prime, ∑' n : ℕ, f (p ^ n) = ∏ p ∈ s with p.Prime, (1 - f p)⁻¹ := by refine prod_congr rfl fun p hp ↦ ?_ simp only [map_pow] exact tsum_geometric_of_norm_lt_one <| h (mem_filter.mp hp).2 @@ -310,7 +310,7 @@ lemma summable_and_hasSum_factoredNumbers_prod_filter_prime_geometric {f : ℕ in terms of the value of the series. -/ lemma prod_filter_prime_geometric_eq_tsum_factoredNumbers {f : ℕ →* F} (hsum : Summable f) (s : Finset ℕ) : - ∏ p ∈ s.filter Nat.Prime, (1 - f p)⁻¹ = ∑' m : factoredNumbers s, f m := by + ∏ p ∈ s with p.Prime, (1 - f p)⁻¹ = ∑' m : factoredNumbers s, f m := by refine (summable_and_hasSum_factoredNumbers_prod_filter_prime_geometric ?_ s).2.tsum_eq.symm exact fun {_} hp ↦ hsum.norm_lt_one hp.one_lt diff --git a/Mathlib/NumberTheory/JacobiSum/Basic.lean b/Mathlib/NumberTheory/JacobiSum/Basic.lean index 7e6fd3e597e49..2de42f2e7d95b 100644 --- a/Mathlib/NumberTheory/JacobiSum/Basic.lean +++ b/Mathlib/NumberTheory/JacobiSum/Basic.lean @@ -107,7 +107,7 @@ theorem jacobiSum_trivial_trivial : simpa only [isUnit_iff_ne_zero, mul_ne_zero_iff, ne_eq, sub_eq_zero, @eq_comm _ _ x] using hx calc ∑ x ∈ univ \ {0, 1}, (MulChar.trivial F R) x * (MulChar.trivial F R) (1 - x) _ = ∑ _ ∈ univ \ {0, 1}, 1 := sum_congr rfl this - _ = Finset.card (univ \ {0, 1}) := (cast_card _).symm + _ = #(univ \ {0, 1}) := (cast_card _).symm _ = Fintype.card F - 2 := by rw [card_sdiff (subset_univ _), card_univ, card_pair zero_ne_one, Nat.cast_sub <| Nat.add_one_le_of_lt Fintype.one_lt_card, Nat.cast_two] diff --git a/Mathlib/NumberTheory/LegendreSymbol/GaussEisensteinLemmas.lean b/Mathlib/NumberTheory/LegendreSymbol/GaussEisensteinLemmas.lean index edf39ca7e1be8..6927e41b5eeef 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/GaussEisensteinLemmas.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/GaussEisensteinLemmas.lean @@ -57,10 +57,9 @@ theorem Ico_map_valMinAbs_natAbs_eq_Ico_map_id (p : ℕ) [hp : Fact p.Prime] (a (fun x _ => (a * x : ZMod p).valMinAbs.natAbs) hmem (inj_on_of_surj_on_of_card_le _ hmem hsurj le_rfl) hsurj (fun _ _ => rfl) -private theorem gauss_lemma_aux₁ (p : ℕ) [Fact p.Prime] {a : ℤ} - (hap : (a : ZMod p) ≠ 0) : (a ^ (p / 2) * (p / 2)! : ZMod p) = - (-1 : ZMod p) ^ ((Ico 1 (p / 2).succ).filter fun x : ℕ => - ¬(a * x : ZMod p).val ≤ p / 2).card * (p / 2)! := +private theorem gauss_lemma_aux₁ (p : ℕ) [Fact p.Prime] {a : ℤ} (hap : (a : ZMod p) ≠ 0) : + (a ^ (p / 2) * (p / 2)! : ZMod p) = + (-1 : ZMod p) ^ #{x ∈ Ico 1 (p / 2).succ | ¬ (a * x.cast : ZMod p).val ≤ p / 2} * (p / 2)! := calc (a ^ (p / 2) * (p / 2)! : ZMod p) = ∏ x ∈ Ico 1 (p / 2).succ, a * x := by rw [prod_mul_distrib, ← prod_natCast, prod_Ico_id_eq_factorial, prod_const, card_Ico, @@ -71,26 +70,25 @@ private theorem gauss_lemma_aux₁ (p : ℕ) [Fact p.Prime] {a : ℤ} (prod_congr rfl fun _ _ => by simp only [natCast_natAbs_valMinAbs] split_ifs <;> simp) - _ = (-1 : ZMod p) ^ ((Ico 1 (p / 2).succ).filter fun x : ℕ => - ¬(a * x : ZMod p).val ≤ p / 2).card * ∏ x ∈ Ico 1 (p / 2).succ, - ↑((a * x : ZMod p).valMinAbs.natAbs) := by + _ = (-1 : ZMod p) ^ #{x ∈ Ico 1 (p / 2).succ | ¬(a * x.cast : ZMod p).val ≤ p / 2} * + ∏ x ∈ Ico 1 (p / 2).succ, ↑((a * x : ZMod p).valMinAbs.natAbs) := by have : (∏ x ∈ Ico 1 (p / 2).succ, if (a * x : ZMod p).val ≤ p / 2 then (1 : ZMod p) else -1) = - ∏ x ∈ (Ico 1 (p / 2).succ).filter fun x : ℕ => ¬(a * x : ZMod p).val ≤ p / 2, -1 := + ∏ x ∈ Ico 1 (p / 2).succ with ¬(a * x.cast : ZMod p).val ≤ p / 2, -1 := prod_bij_ne_one (fun x _ _ => x) (fun x => by split_ifs <;> (dsimp; simp_all)) (fun _ _ _ _ _ _ => id) (fun b h _ => ⟨b, by simp_all [-not_le]⟩) (by intros; split_ifs at * <;> simp_all) rw [prod_mul_distrib, this, prod_const] - _ = (-1 : ZMod p) ^ ((Ico 1 (p / 2).succ).filter fun x : ℕ => - ¬(a * x : ZMod p).val ≤ p / 2).card * (p / 2)! := by + _ = (-1 : ZMod p) ^ #{x ∈ Ico 1 (p / 2).succ | ¬(a * x.cast : ZMod p).val ≤ p / 2} * + (p / 2)! := by rw [← prod_natCast, Finset.prod_eq_multiset_prod, Ico_map_valMinAbs_natAbs_eq_Ico_map_id p a hap, ← Finset.prod_eq_multiset_prod, prod_Ico_id_eq_factorial] -theorem gauss_lemma_aux (p : ℕ) [hp : Fact p.Prime] {a : ℤ} - (hap : (a : ZMod p) ≠ 0) : (↑a ^ (p / 2) : ZMod p) = - ((-1) ^ ((Ico 1 (p / 2).succ).filter fun x : ℕ => p / 2 < (a * x : ZMod p).val).card :) := +theorem gauss_lemma_aux (p : ℕ) [hp : Fact p.Prime] {a : ℤ} (hap : (a : ZMod p) ≠ 0) : + (a ^ (p / 2) : ZMod p) = + ((-1) ^ #{x ∈ Ico 1 (p / 2).succ | p / 2 < (a * x.cast : ZMod p).val} :) := (mul_left_inj' (show ((p / 2)! : ZMod p) ≠ 0 by rw [Ne, CharP.cast_eq_zero_iff (ZMod p) p, hp.1.dvd_factorial, not_le] exact Nat.div_lt_self hp.1.pos (by decide))).1 <| by @@ -99,21 +97,20 @@ theorem gauss_lemma_aux (p : ℕ) [hp : Fact p.Prime] {a : ℤ} /-- **Gauss' lemma**. The Legendre symbol can be computed by considering the number of naturals less than `p/2` such that `(a * x) % p > p / 2`. -/ theorem gauss_lemma {p : ℕ} [h : Fact p.Prime] {a : ℤ} (hp : p ≠ 2) (ha0 : (a : ZMod p) ≠ 0) : - legendreSym p a = (-1) ^ ((Ico 1 (p / 2).succ).filter fun x : ℕ => - p / 2 < (a * x : ZMod p).val).card := by + legendreSym p a = (-1) ^ #{x ∈ Ico 1 (p / 2).succ | p / 2 < (a * x.cast : ZMod p).val} := by replace hp : Odd p := h.out.odd_of_ne_two hp - have : (legendreSym p a : ZMod p) = (((-1) ^ ((Ico 1 (p / 2).succ).filter fun x : ℕ => - p / 2 < (a * x : ZMod p).val).card : ℤ) : ZMod p) := by + have : (legendreSym p a : ZMod p) = + (((-1) ^ #{x ∈ Ico 1 (p / 2).succ | p / 2 < (a * x.cast : ZMod p).val} : ℤ) : ZMod p) := by rw [legendreSym.eq_pow, gauss_lemma_aux p ha0] cases legendreSym.eq_one_or_neg_one p ha0 <;> - cases neg_one_pow_eq_or ℤ - ((Ico 1 (p / 2).succ).filter fun x : ℕ => p / 2 < (a * x : ZMod p).val).card <;> + cases neg_one_pow_eq_or ℤ #{x ∈ Ico 1 (p / 2).succ | p / 2 < (a * x.cast : ZMod p).val} <;> simp_all [ne_neg_self hp one_ne_zero, (ne_neg_self hp one_ne_zero).symm] private theorem eisenstein_lemma_aux₁ (p : ℕ) [Fact p.Prime] [hp2 : Fact (p % 2 = 1)] {a : ℕ} - (hap : (a : ZMod p) ≠ 0) : ((∑ x ∈ Ico 1 (p / 2).succ, a * x : ℕ) : ZMod 2) = - ((Ico 1 (p / 2).succ).filter fun x : ℕ => p / 2 < (a * x : ZMod p).val).card + - ∑ x ∈ Ico 1 (p / 2).succ, x + (∑ x ∈ Ico 1 (p / 2).succ, a * x / p : ℕ) := + (hap : (a : ZMod p) ≠ 0) : + ((∑ x ∈ Ico 1 (p / 2).succ, a * x : ℕ) : ZMod 2) = + #{x ∈ Ico 1 (p / 2).succ | p / 2 < (a * x.cast : ZMod p).val} + + ∑ x ∈ Ico 1 (p / 2).succ, x + (∑ x ∈ Ico 1 (p / 2).succ, a * x / p : ℕ) := have hp2 : (p : ZMod 2) = (1 : ℕ) := (eq_iff_modEq_nat _).2 hp2.1 calc ((∑ x ∈ Ico 1 (p / 2).succ, a * x : ℕ) : ZMod 2) = @@ -130,8 +127,8 @@ private theorem eisenstein_lemma_aux₁ (p : ℕ) [Fact p.Prime] [hp2 : Fact (p ∑ x ∈ Ico 1 (p / 2).succ, (((a * x : ZMod p).valMinAbs + if (a * x : ZMod p).val ≤ p / 2 then 0 else p : ℤ) : ZMod 2) := by simp only [(val_eq_ite_valMinAbs _).symm]; simp [Nat.cast_sum] - _ = ((Ico 1 (p / 2).succ).filter fun x : ℕ => p / 2 < (a * x : ZMod p).val).card + - (∑ x ∈ Ico 1 (p / 2).succ, (a * x : ZMod p).valMinAbs.natAbs : ℕ) := by + _ = #{x ∈ Ico 1 (p / 2).succ | p / 2 < (a * x.cast : ZMod p).val} + + (∑ x ∈ Ico 1 (p / 2).succ, (a * x.cast : ZMod p).valMinAbs.natAbs : ℕ) := by simp [add_comm, sum_add_distrib, Finset.sum_ite, hp2, Nat.cast_sum] _ = _ := by rw [Finset.sum_eq_multiset_sum, Ico_map_valMinAbs_natAbs_eq_Ico_map_id p a hap, ← @@ -140,7 +137,7 @@ private theorem eisenstein_lemma_aux₁ (p : ℕ) [Fact p.Prime] [hp2 : Fact (p theorem eisenstein_lemma_aux (p : ℕ) [Fact p.Prime] [Fact (p % 2 = 1)] {a : ℕ} (ha2 : a % 2 = 1) (hap : (a : ZMod p) ≠ 0) : - ((Ico 1 (p / 2).succ).filter fun x : ℕ => p / 2 < (a * x : ZMod p).val).card ≡ + #{x ∈ Ico 1 (p / 2).succ | p / 2 < (a * x.cast : ZMod p).val} ≡ ∑ x ∈ Ico 1 (p / 2).succ, x * a / p [MOD 2] := have ha2 : (a : ZMod 2) = (1 : ℕ) := (eq_iff_modEq_nat _).2 ha2 (eq_iff_modEq_nat 2).1 <| sub_eq_zero.1 <| by @@ -149,10 +146,10 @@ theorem eisenstein_lemma_aux (p : ℕ) [Fact p.Prime] [Fact (p % 2 = 1)] {a : Eq.symm (eisenstein_lemma_aux₁ p hap) theorem div_eq_filter_card {a b c : ℕ} (hb0 : 0 < b) (hc : a / b ≤ c) : - a / b = ((Ico 1 c.succ).filter fun x => x * b ≤ a).card := + a / b = #{x ∈ Ico 1 c.succ | x * b ≤ a} := calc - a / b = (Ico 1 (a / b).succ).card := by simp - _ = ((Ico 1 c.succ).filter fun x => x * b ≤ a).card := + a / b = #(Ico 1 (a / b).succ) := by simp + _ = #{x ∈ Ico 1 c.succ | x * b ≤ a} := congr_arg _ <| Finset.ext fun x => by have : x * b ≤ a → x ≤ c := fun h => le_trans (by rwa [le_div_iff_mul_le hb0]) hc simp [Nat.lt_succ_iff, le_div_iff_mul_le hb0]; tauto @@ -161,13 +158,12 @@ theorem div_eq_filter_card {a b c : ℕ} (hb0 : 0 < b) (hc : a / b ≤ c) : rectangle `(0, p/2) × (0, q/2)`. -/ private theorem sum_Ico_eq_card_lt {p q : ℕ} : ∑ a ∈ Ico 1 (p / 2).succ, a * q / p = - ((Ico 1 (p / 2).succ ×ˢ Ico 1 (q / 2).succ).filter fun x : ℕ × ℕ => - x.2 * p ≤ x.1 * q).card := + #{x ∈ Ico 1 (p / 2).succ ×ˢ Ico 1 (q / 2).succ | x.2 * p ≤ x.1 * q} := if hp0 : p = 0 then by simp [hp0, Finset.ext_iff] else calc ∑ a ∈ Ico 1 (p / 2).succ, a * q / p = - ∑ a ∈ Ico 1 (p / 2).succ, ((Ico 1 (q / 2).succ).filter fun x => x * p ≤ a * q).card := + ∑ a ∈ Ico 1 (p / 2).succ, #{x ∈ Ico 1 (q / 2).succ | x * p ≤ a * q} := Finset.sum_congr rfl fun x hx => div_eq_filter_card (Nat.pos_of_ne_zero hp0) <| calc x * q / p ≤ p / 2 * q / p := by have := le_of_lt_succ (mem_Ico.mp hx).2; gcongr @@ -187,16 +183,15 @@ theorem sum_mul_div_add_sum_mul_div_eq_mul (p q : ℕ) [hp : Fact p.Prime] (hq0 ∑ a ∈ Ico 1 (p / 2).succ, a * q / p + ∑ a ∈ Ico 1 (q / 2).succ, a * p / q = p / 2 * (q / 2) := by have hswap : - ((Ico 1 (q / 2).succ ×ˢ Ico 1 (p / 2).succ).filter fun x : ℕ × ℕ => x.2 * q ≤ x.1 * p).card = - ((Ico 1 (p / 2).succ ×ˢ Ico 1 (q / 2).succ).filter fun x : ℕ × ℕ => - x.1 * q ≤ x.2 * p).card := + #{x ∈ Ico 1 (q / 2).succ ×ˢ Ico 1 (p / 2).succ | x.2 * q ≤ x.1 * p} = + #{x ∈ Ico 1 (p / 2).succ ×ˢ Ico 1 (q / 2).succ | x.1 * q ≤ x.2 * p} := 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, 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 + Disjoint {x ∈ Ico 1 (p / 2).succ ×ˢ Ico 1 (q / 2).succ | x.2 * p ≤ x.1 * q} + {x ∈ Ico 1 (p / 2).succ ×ˢ Ico 1 (q / 2).succ | x.1 * q ≤ x.2 * p} := by apply disjoint_filter.2 fun x hx hpq hqp => ?_ have hxp : x.1 < p := lt_of_le_of_lt (show x.1 ≤ p / 2 by simp_all only [Nat.lt_succ_iff, mem_Ico, mem_product]) @@ -207,8 +202,8 @@ theorem sum_mul_div_add_sum_mul_div_eq_mul (p q : ℕ) [hp : Fact p.Prime] (hq0 rw [val_cast_of_lt hxp, val_zero] at this simp only [this, nonpos_iff_eq_zero, mem_Ico, one_ne_zero, false_and, mem_product] at hx have hunion : - (((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) = + {x ∈ Ico 1 (p / 2).succ ×ˢ Ico 1 (q / 2).succ | x.2 * p ≤ x.1 * q} ∪ + {x ∈ Ico 1 (p / 2).succ ×ˢ Ico 1 (q / 2).succ | x.1 * q ≤ x.2 * p} = Ico 1 (p / 2).succ ×ˢ Ico 1 (q / 2).succ := Finset.ext fun x => by have := le_total (x.2 * p) (x.1 * q) diff --git a/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean b/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean index 9ba104ca9690a..a46bc6602b64d 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean @@ -223,7 +223,7 @@ theorem quadraticChar_isNontrivial (hF : ringChar F ≠ 2) : (quadraticChar F).I open Finset in /-- The number of solutions to `x^2 = a` is determined by the quadratic character. -/ theorem quadraticChar_card_sqrts (hF : ringChar F ≠ 2) (a : F) : - ↑{x : F | x ^ 2 = a}.toFinset.card = quadraticChar F a + 1 := by + #{x : F | x ^ 2 = a}.toFinset = quadraticChar F a + 1 := by -- we consider the cases `a = 0`, `a` is a nonzero square and `a` is a nonsquare in turn by_cases h₀ : a = 0 · simp only [h₀, sq_eq_zero_iff, Set.setOf_eq_eq_singleton, Set.toFinset_card, diff --git a/Mathlib/NumberTheory/MaricaSchoenheim.lean b/Mathlib/NumberTheory/MaricaSchoenheim.lean index 4ce9d2d8bf921..a88458695d67d 100644 --- a/Mathlib/NumberTheory/MaricaSchoenheim.lean +++ b/Mathlib/NumberTheory/MaricaSchoenheim.lean @@ -40,9 +40,9 @@ lemma grahamConjecture_of_squarefree {n : ℕ} (f : ℕ → ℕ) (hf' : ∀ k < fun i hi j ↦ (hf' _ hi).squarefree_of_dvd <| div_dvd_of_dvd <| gcd_dvd_left _ _ refine lt_irrefl n ?_ calc - n = 𝒜.card := ?_ - _ ≤ (𝒜 \\ 𝒜).card := 𝒜.card_le_card_diffs - _ ≤ (Ioo 0 n).card := card_le_card_of_injOn (fun s ↦ ∏ p ∈ s, p) ?_ ?_ + n = #𝒜 := ?_ + _ ≤ #(𝒜 \\ 𝒜) := 𝒜.card_le_card_diffs + _ ≤ #(Ioo 0 n) := card_le_card_of_injOn (fun s ↦ ∏ p ∈ s, p) ?_ ?_ _ = n - 1 := by rw [card_Ioo, tsub_zero] _ < n := tsub_lt_self hn.bot_lt zero_lt_one · rw [Finset.card_image_of_injOn, card_Iio] diff --git a/Mathlib/NumberTheory/NumberField/Embeddings.lean b/Mathlib/NumberTheory/NumberField/Embeddings.lean index d713d6bce1442..d5e8af77ec71d 100644 --- a/Mathlib/NumberTheory/NumberField/Embeddings.lean +++ b/Mathlib/NumberTheory/NumberField/Embeddings.lean @@ -32,7 +32,7 @@ for `x ∈ K`, we have `Π_w ‖x‖_w = |norm(x)|` where the product is over th number field, embeddings, places, infinite places -/ -open scoped Classical +open scoped Classical Finset namespace NumberField.Embeddings @@ -432,8 +432,7 @@ theorem one_le_mult {w : InfinitePlace K} : (1 : ℝ) ≤ mult w := by rw [← Nat.cast_one, Nat.cast_le] exact mult_pos -theorem card_filter_mk_eq [NumberField K] (w : InfinitePlace K) : - (Finset.univ.filter fun φ => mk φ = w).card = mult w := by +theorem card_filter_mk_eq [NumberField K] (w : InfinitePlace K) : #{φ | mk φ = w} = mult w := by conv_lhs => congr; congr; ext rw [← mk_embedding w, mk_eq_iff, ComplexEmbedding.conjugate, star_involutive.eq_iff] @@ -490,11 +489,8 @@ theorem prod_eq_abs_norm (x : K) : · rw [map_prod, ← Fintype.prod_equiv RingHom.equivRatAlgHom (fun f => Complex.abs (f x)) (fun φ => Complex.abs (φ x)) fun _ => by simp [RingHom.equivRatAlgHom_apply]; rfl] rw [← Finset.prod_fiberwise Finset.univ mk (fun φ => Complex.abs (φ x))] - have : ∀ w : InfinitePlace K, ∀ φ ∈ Finset.filter (fun a ↦ mk a = w) Finset.univ, - Complex.abs (φ x) = w x := by - intro _ _ hφ - rw [← (Finset.mem_filter.mp hφ).2] - rfl + have (w : InfinitePlace K) (φ) (hφ : φ ∈ ({φ | mk φ = w} : Finset _)) : + Complex.abs (φ x) = w x := by rw [← (Finset.mem_filter.mp hφ).2, apply] simp_rw [Finset.prod_congr rfl (this _), Finset.prod_const, card_filter_mk_eq] · rw [eq_ratCast, Rat.cast_abs, ← Complex.abs_ofReal, Complex.ofReal_ratCast] @@ -571,8 +567,8 @@ theorem card_eq_nrRealPlaces_add_nrComplexPlaces : theorem card_complex_embeddings : card { φ : K →+* ℂ // ¬ComplexEmbedding.IsReal φ } = 2 * NrComplexPlaces K := by - suffices ∀ w : { w : InfinitePlace K // IsComplex w }, (Finset.univ.filter - fun φ : { φ // ¬ ComplexEmbedding.IsReal φ } => mkComplex φ = w).card = 2 by + suffices ∀ w : { w : InfinitePlace K // IsComplex w }, + #{φ : {φ //¬ ComplexEmbedding.IsReal φ} | mkComplex φ = w} = 2 by rw [Fintype.card, Finset.card_eq_sum_ones, ← Finset.sum_fiberwise _ (fun φ => mkComplex φ)] simp_rw [Finset.sum_const, this, smul_eq_mul, mul_one, Fintype.card, Finset.card_eq_sum_ones, Finset.mul_sum, Finset.sum_const, smul_eq_mul, mul_one] @@ -930,16 +926,16 @@ variable [NumberField K] open Finset in lemma card_isUnramified [NumberField k] [IsGalois k K] : - Finset.card (univ.filter <| IsUnramified k (K := K)) = - Finset.card (univ.filter <| IsUnramifiedIn K (k := k)) * (finrank k K) := by + #{w : InfinitePlace K | w.IsUnramified k} = + #{w : InfinitePlace k | w.IsUnramifiedIn K} * finrank k K := by letI := Module.Finite.of_restrictScalars_finite ℚ k K rw [← IsGalois.card_aut_eq_finrank, Finset.card_eq_sum_card_fiberwise (f := (comap · (algebraMap k K))) - (t := (univ.filter <| IsUnramifiedIn K (k := k))), ← smul_eq_mul, ← sum_const] + (t := {w : InfinitePlace k | w.IsUnramifiedIn K}), ← smul_eq_mul, ← sum_const] · refine sum_congr rfl (fun w hw ↦ ?_) obtain ⟨w, rfl⟩ := comap_surjective (K := K) w simp only [mem_univ, forall_true_left, mem_filter, true_and] at hw - trans Finset.card (MulAction.orbit (K ≃ₐ[k] K) w).toFinset + trans #(MulAction.orbit (K ≃ₐ[k] K) w).toFinset · congr; ext w' simp only [mem_univ, forall_true_left, filter_congr_decidable, mem_filter, true_and, Set.mem_toFinset, mem_orbit_iff, @eq_comm _ (comap w' _), and_iff_right_iff_imp] @@ -952,12 +948,12 @@ lemma card_isUnramified [NumberField k] [IsGalois k K] : open Finset in lemma card_isUnramified_compl [NumberField k] [IsGalois k K] : - Finset.card (univ.filter <| IsUnramified k (K := K))ᶜ = - Finset.card (univ.filter <| IsUnramifiedIn K (k := k))ᶜ * (finrank k K / 2) := by + #({w : InfinitePlace K | w.IsUnramified k} : Finset _)ᶜ = + #({w : InfinitePlace k | w.IsUnramifiedIn K} : Finset _)ᶜ * (finrank k K / 2) := by letI := Module.Finite.of_restrictScalars_finite ℚ k K rw [← IsGalois.card_aut_eq_finrank, Finset.card_eq_sum_card_fiberwise (f := (comap · (algebraMap k K))) - (t := (univ.filter <| IsUnramifiedIn K (k := k))ᶜ), ← smul_eq_mul, ← sum_const] + (t := ({w : InfinitePlace k | w.IsUnramifiedIn K}: Finset _)ᶜ), ← smul_eq_mul, ← sum_const] · refine sum_congr rfl (fun w hw ↦ ?_) obtain ⟨w, rfl⟩ := comap_surjective (K := K) w simp only [mem_univ, forall_true_left, compl_filter, not_not, mem_filter, true_and] at hw @@ -974,8 +970,8 @@ lemma card_isUnramified_compl [NumberField k] [IsGalois k K] : lemma card_eq_card_isUnramifiedIn [NumberField k] [IsGalois k K] : Fintype.card (InfinitePlace K) = - Finset.card (Finset.univ.filter <| IsUnramifiedIn K (k := k)) * finrank k K + - Finset.card (Finset.univ.filter <| IsUnramifiedIn K (k := k))ᶜ * (finrank k K / 2) := by + #{w : InfinitePlace k | w.IsUnramifiedIn K} * finrank k K + + #({w : InfinitePlace k | w.IsUnramifiedIn K} : Finset _)ᶜ * (finrank k K / 2) := by rw [← card_isUnramified, ← card_isUnramified_compl, Finset.card_add_card_compl] end NumberField.InfinitePlace diff --git a/Mathlib/NumberTheory/NumberField/EquivReindex.lean b/Mathlib/NumberTheory/NumberField/EquivReindex.lean index 0226ce9f8e5d5..76ece2db3d216 100644 --- a/Mathlib/NumberTheory/NumberField/EquivReindex.lean +++ b/Mathlib/NumberTheory/NumberField/EquivReindex.lean @@ -11,8 +11,9 @@ import Mathlib.NumberTheory.NumberField.CanonicalEmbedding.Basic This file introduces an equivalence between the set of embeddings of `K` into `ℂ` and the index set of the chosen basis of the ring of integers of `K`. -## Tagshouse -number field, algebraic number +## Tags + +house, number field, algebraic number -/ variable (K : Type*) [Field K] [NumberField K] diff --git a/Mathlib/NumberTheory/PrimeCounting.lean b/Mathlib/NumberTheory/PrimeCounting.lean index 3005d0cb7165a..a488afee078ab 100644 --- a/Mathlib/NumberTheory/PrimeCounting.lean +++ b/Mathlib/NumberTheory/PrimeCounting.lean @@ -101,7 +101,7 @@ theorem prime_nth_prime (n : ℕ) : Prime (nth Prime n) := /-- The cardinality of the finset `primesBelow n` equals the counting function `primeCounting'` at `n`. -/ -theorem primesBelow_card_eq_primeCounting' (n : ℕ) : n.primesBelow.card = primeCounting' n := by +theorem primesBelow_card_eq_primeCounting' (n : ℕ) : #n.primesBelow = primeCounting' n := by simp only [primesBelow, primeCounting'] exact (count_eq_card_filter_range Prime n).symm @@ -109,13 +109,13 @@ theorem primesBelow_card_eq_primeCounting' (n : ℕ) : n.primesBelow.card = prim theorem primeCounting'_add_le {a k : ℕ} (h0 : 0 < a) (h1 : a < k) (n : ℕ) : π' (k + n) ≤ π' k + Nat.totient a * (n / a + 1) := calc - π' (k + n) ≤ ((range k).filter Prime).card + ((Ico k (k + n)).filter Prime).card := by + π' (k + n) ≤ #{p ∈ range k | p.Prime} + #{p ∈ Ico k (k + n) | p.Prime} := by rw [primeCounting', count_eq_card_filter_range, range_eq_Ico, ← Ico_union_Ico_eq_Ico (zero_le k) le_self_add, filter_union] apply card_union_le - _ ≤ π' k + ((Ico k (k + n)).filter Prime).card := by + _ ≤ π' k + #{p ∈ Ico k (k + n) | p.Prime} := by rw [primeCounting', count_eq_card_filter_range] - _ ≤ π' k + ((Ico k (k + n)).filter (Coprime a)).card := by + _ ≤ π' k + #{b ∈ Ico k (k + n) | a.Coprime b} := by refine add_le_add_left (card_le_card ?_) k.primeCounting' simp only [subset_iff, and_imp, mem_filter, mem_Ico] intro p succ_k_le_p p_lt_n p_prime diff --git a/Mathlib/NumberTheory/Primorial.lean b/Mathlib/NumberTheory/Primorial.lean index 886f7183cfa38..7ef909c7b0799 100644 --- a/Mathlib/NumberTheory/Primorial.lean +++ b/Mathlib/NumberTheory/Primorial.lean @@ -31,8 +31,7 @@ open Nat /-- The primorial `n#` of `n` is the product of the primes less than or equal to `n`. -/ -def primorial (n : ℕ) : ℕ := - ∏ p ∈ filter Nat.Prime (range (n + 1)), p +def primorial (n : ℕ) : ℕ := ∏ p ∈ range (n + 1) with p.Prime, p local notation x "#" => primorial x @@ -45,14 +44,14 @@ theorem primorial_succ {n : ℕ} (hn1 : n ≠ 1) (hn : Odd n) : (n + 1)# = n# := exact fun h ↦ h.even_sub_one <| mt succ.inj hn1 theorem primorial_add (m n : ℕ) : - (m + n)# = m# * ∏ p ∈ filter Nat.Prime (Ico (m + 1) (m + n + 1)), p := by + (m + n)# = m# * ∏ p ∈ Ico (m + 1) (m + n + 1) with p.Prime, p := by rw [primorial, primorial, ← Ico_zero_eq_range, ← prod_union, ← filter_union, Ico_union_Ico_eq_Ico] exacts [Nat.zero_le _, add_le_add_right (Nat.le_add_right _ _) _, disjoint_filter_filter <| Ico_disjoint_Ico_consecutive _ _ _] theorem primorial_add_dvd {m n : ℕ} (h : n ≤ m) : (m + n)# ∣ m# * choose (m + n) m := calc - (m + n)# = m# * ∏ p ∈ filter Nat.Prime (Ico (m + 1) (m + n + 1)), p := primorial_add _ _ + (m + n)# = m# * ∏ p ∈ Ico (m + 1) (m + n + 1) with p.Prime, p := primorial_add _ _ _ ∣ m# * choose (m + n) m := mul_dvd_mul_left _ <| prod_primes_dvd _ (fun _ hk ↦ (mem_filter.1 hk).2.prime) fun p hp ↦ by diff --git a/Mathlib/NumberTheory/SiegelsLemma.lean b/Mathlib/NumberTheory/SiegelsLemma.lean index 998f8940d66b9..22160797bf7a5 100644 --- a/Mathlib/NumberTheory/SiegelsLemma.lean +++ b/Mathlib/NumberTheory/SiegelsLemma.lean @@ -61,7 +61,7 @@ section preparation /- In order to apply Pigeonhole we need: # Step 1: ∀ v ∈ T, A *ᵥ v ∈ S and -# Step 2: S.card < T.card +# Step 2: #S < #T Pigeonhole will give different x and y in T with A.mulVec x = A.mulVec y in S Their difference is the solution we are looking for -/ @@ -91,12 +91,12 @@ private lemma image_T_subset_S [DecidableEq α] [DecidableEq β] (v) (hv : v ∈ -- # Preparation for Step 2 -private lemma card_T_eq [DecidableEq β] : (T).card = (B + 1) ^ n := by +private lemma card_T_eq [DecidableEq β] : #T = (B + 1) ^ n := by rw [Pi.card_Icc 0 B'] simp only [Pi.zero_apply, card_Icc, sub_zero, toNat_ofNat_add_one, prod_const, card_univ, add_pos_iff, zero_lt_one, or_true] --- This lemma is necessary to be able to apply the formula (Icc a b).card = b + 1 - a +-- This lemma is necessary to be able to apply the formula #(Icc a b) = b + 1 - a private lemma N_le_P_add_one (i : α) : N i ≤ P i + 1 := by calc N i _ ≤ 0 := by @@ -109,7 +109,7 @@ private lemma N_le_P_add_one (i : α) : N i ≤ P i + 1 := by intro j _ exact mul_nonneg (Nat.cast_nonneg B) (posPart_nonneg (A i j)) -private lemma card_S_eq [DecidableEq α] : (Finset.Icc N P).card = ∏ i : α, (P i - N i + 1) := by +private lemma card_S_eq [DecidableEq α] : #(Finset.Icc N P) = ∏ i : α, (P i - N i + 1) := by rw [Pi.card_Icc N P, Nat.cast_prod] congr ext i @@ -128,13 +128,13 @@ lemma one_le_norm_A_of_ne_zero (hA : A ≠ 0) : 1 ≤ ‖A‖ := by norm_cast at h exact Int.abs_lt_one_iff.1 h --- # Step 2: S.card < T.card +-- # Step 2: #S < #T open Real Nat private lemma card_S_lt_card_T [DecidableEq α] [DecidableEq β] (hn : Fintype.card α < Fintype.card β) (hm : 0 < Fintype.card α) : - (S).card < (T).card := by + #S < #T := by zify -- This is necessary to use card_S_eq rw [card_T_eq A, card_S_eq] rify -- This is necessary because ‖A‖ is a real number diff --git a/Mathlib/Order/Interval/Finset/Basic.lean b/Mathlib/Order/Interval/Finset/Basic.lean index 53ccb40f79265..4f194b2424862 100644 --- a/Mathlib/Order/Interval/Finset/Basic.lean +++ b/Mathlib/Order/Interval/Finset/Basic.lean @@ -269,66 +269,58 @@ def _root_.Set.fintypeOfMemBounds {s : Set α} [DecidablePred (· ∈ s)] (ha : section Filter theorem Ico_filter_lt_of_le_left [DecidablePred (· < c)] (hca : c ≤ a) : - (Ico a b).filter (· < c) = ∅ := + {x ∈ Ico a b | x < c} = ∅ := filter_false_of_mem fun _ hx => (hca.trans (mem_Ico.1 hx).1).not_lt theorem Ico_filter_lt_of_right_le [DecidablePred (· < c)] (hbc : b ≤ c) : - (Ico a b).filter (· < c) = Ico a b := + {x ∈ Ico a b | x < c} = Ico a b := filter_true_of_mem fun _ hx => (mem_Ico.1 hx).2.trans_le hbc theorem Ico_filter_lt_of_le_right [DecidablePred (· < c)] (hcb : c ≤ b) : - (Ico a b).filter (· < c) = Ico a c := by + {x ∈ Ico a b | x < c} = Ico a c := by ext x rw [mem_filter, mem_Ico, mem_Ico, and_right_comm] exact and_iff_left_of_imp fun h => h.2.trans_le hcb theorem Ico_filter_le_of_le_left {a b c : α} [DecidablePred (c ≤ ·)] (hca : c ≤ a) : - (Ico a b).filter (c ≤ ·) = Ico a b := + {x ∈ Ico a b | c ≤ x} = Ico a b := filter_true_of_mem fun _ hx => hca.trans (mem_Ico.1 hx).1 theorem Ico_filter_le_of_right_le {a b : α} [DecidablePred (b ≤ ·)] : - (Ico a b).filter (b ≤ ·) = ∅ := + {x ∈ Ico a b | b ≤ x} = ∅ := filter_false_of_mem fun _ hx => (mem_Ico.1 hx).2.not_le theorem Ico_filter_le_of_left_le {a b c : α} [DecidablePred (c ≤ ·)] (hac : a ≤ c) : - (Ico a b).filter (c ≤ ·) = Ico c b := by + {x ∈ Ico a b | c ≤ x} = Ico c b := by ext x rw [mem_filter, mem_Ico, mem_Ico, and_comm, and_left_comm] exact and_iff_right_of_imp fun h => hac.trans h.1 theorem Icc_filter_lt_of_lt_right {a b c : α} [DecidablePred (· < c)] (h : b < c) : - (Icc a b).filter (· < c) = Icc a b := + {x ∈ Icc a b | x < c} = Icc a b := filter_true_of_mem fun _ hx => lt_of_le_of_lt (mem_Icc.1 hx).2 h theorem Ioc_filter_lt_of_lt_right {a b c : α} [DecidablePred (· < c)] (h : b < c) : - (Ioc a b).filter (· < c) = Ioc a b := + {x ∈ Ioc a b | x < c} = Ioc a b := filter_true_of_mem fun _ hx => lt_of_le_of_lt (mem_Ioc.1 hx).2 h theorem Iic_filter_lt_of_lt_right {α} [Preorder α] [LocallyFiniteOrderBot α] {a c : α} - [DecidablePred (· < c)] (h : a < c) : (Iic a).filter (· < c) = Iic a := + [DecidablePred (· < c)] (h : a < c) : {x ∈ Iic a | x < c} = Iic a := filter_true_of_mem fun _ hx => lt_of_le_of_lt (mem_Iic.1 hx) h variable (a b) [Fintype α] theorem filter_lt_lt_eq_Ioo [DecidablePred fun j => a < j ∧ j < b] : - (univ.filter fun j => a < j ∧ j < b) = Ioo a b := by - ext - simp + ({j | a < j ∧ j < b} : Finset _) = Ioo a b := by ext; simp theorem filter_lt_le_eq_Ioc [DecidablePred fun j => a < j ∧ j ≤ b] : - (univ.filter fun j => a < j ∧ j ≤ b) = Ioc a b := by - ext - simp + ({j | a < j ∧ j ≤ b} : Finset _) = Ioc a b := by ext; simp theorem filter_le_lt_eq_Ico [DecidablePred fun j => a ≤ j ∧ j < b] : - (univ.filter fun j => a ≤ j ∧ j < b) = Ico a b := by - ext - simp + ({j | a ≤ j ∧ j < b} : Finset _) = Ico a b := by ext; simp theorem filter_le_le_eq_Icc [DecidablePred fun j => a ≤ j ∧ j ≤ b] : - (univ.filter fun j => a ≤ j ∧ j ≤ b) = Icc a b := by - ext - simp + ({j | a ≤ j ∧ j ≤ b} : Finset _) = Icc a b := by ext; simp end Filter @@ -430,13 +422,8 @@ theorem _root_.Set.Infinite.not_bddBelow {s : Set α} : s.Infinite → ¬BddBelo variable [Fintype α] -theorem filter_lt_eq_Ioi [DecidablePred (a < ·)] : univ.filter (a < ·) = Ioi a := by - ext - simp - -theorem filter_le_eq_Ici [DecidablePred (a ≤ ·)] : univ.filter (a ≤ ·) = Ici a := by - ext - simp +theorem filter_lt_eq_Ioi [DecidablePred (a < ·)] : ({x | a < x} : Finset _) = Ioi a := by ext; simp +theorem filter_le_eq_Ici [DecidablePred (a ≤ ·)] : ({x | a ≤ x} : Finset _) = Ici a := by ext; simp end LocallyFiniteOrderTop @@ -455,13 +442,8 @@ theorem _root_.Set.Infinite.not_bddAbove {s : Set α} : s.Infinite → ¬BddAbov variable [Fintype α] -theorem filter_gt_eq_Iio [DecidablePred (· < a)] : univ.filter (· < a) = Iio a := by - ext - simp - -theorem filter_ge_eq_Iic [DecidablePred (· ≤ a)] : univ.filter (· ≤ a) = Iic a := by - ext - simp +theorem filter_gt_eq_Iio [DecidablePred (· < a)] : ({x | x < a} : Finset _) = Iio a := by ext; simp +theorem filter_ge_eq_Iic [DecidablePred (· ≤ a)] : ({x | x ≤ a} : Finset _) = Iic a := by ext; simp end LocallyFiniteOrderBot @@ -561,32 +543,32 @@ theorem Ico_eq_cons_Ioo (h : a < b) : Ico a b = (Ioo a b).cons a left_not_mem_Io classical rw [cons_eq_insert, Ioo_insert_left h] theorem Ico_filter_le_left {a b : α} [DecidablePred (· ≤ a)] (hab : a < b) : - ((Ico a b).filter fun x => x ≤ a) = {a} := by + {x ∈ Ico a b | x ≤ a} = {a} := by ext x rw [mem_filter, mem_Ico, mem_singleton, and_right_comm, ← le_antisymm_iff, eq_comm] exact and_iff_left_of_imp fun h => h.le.trans_lt hab -theorem card_Ico_eq_card_Icc_sub_one (a b : α) : (Ico a b).card = (Icc a b).card - 1 := by +theorem card_Ico_eq_card_Icc_sub_one (a b : α) : #(Ico a b) = #(Icc a b) - 1 := by classical by_cases h : a ≤ b · rw [Icc_eq_cons_Ico h, card_cons] exact (Nat.add_sub_cancel _ _).symm · rw [Ico_eq_empty fun h' => h h'.le, Icc_eq_empty h, card_empty, Nat.zero_sub] -theorem card_Ioc_eq_card_Icc_sub_one (a b : α) : (Ioc a b).card = (Icc a b).card - 1 := +theorem card_Ioc_eq_card_Icc_sub_one (a b : α) : #(Ioc a b) = #(Icc a b) - 1 := @card_Ico_eq_card_Icc_sub_one αᵒᵈ _ _ _ _ -theorem card_Ioo_eq_card_Ico_sub_one (a b : α) : (Ioo a b).card = (Ico a b).card - 1 := by +theorem card_Ioo_eq_card_Ico_sub_one (a b : α) : #(Ioo a b) = #(Ico a b) - 1 := by classical by_cases h : a < b · rw [Ico_eq_cons_Ioo h, card_cons] exact (Nat.add_sub_cancel _ _).symm · rw [Ioo_eq_empty h, Ico_eq_empty h, card_empty, Nat.zero_sub] -theorem card_Ioo_eq_card_Ioc_sub_one (a b : α) : (Ioo a b).card = (Ioc a b).card - 1 := +theorem card_Ioo_eq_card_Ioc_sub_one (a b : α) : #(Ioo a b) = #(Ioc a b) - 1 := @card_Ioo_eq_card_Ico_sub_one αᵒᵈ _ _ _ _ -theorem card_Ioo_eq_card_Icc_sub_two (a b : α) : (Ioo a b).card = (Icc a b).card - 2 := by +theorem card_Ioo_eq_card_Icc_sub_two (a b : α) : #(Ioo a b) = #(Icc a b) - 2 := by rw [card_Ioo_eq_card_Ico_sub_one, card_Ico_eq_card_Icc_sub_one] rfl @@ -619,7 +601,7 @@ theorem not_mem_Ioi_self {b : α} : b ∉ Ioi b := fun h => lt_irrefl _ (mem_Ioi theorem Ici_eq_cons_Ioi (a : α) : Ici a = (Ioi a).cons a not_mem_Ioi_self := by classical rw [cons_eq_insert, Ioi_insert] -theorem card_Ioi_eq_card_Ici_sub_one (a : α) : (Ioi a).card = (Ici a).card - 1 := by +theorem card_Ioi_eq_card_Ici_sub_one (a : α) : #(Ioi a) = #(Ici a) - 1 := by rw [Ici_eq_cons_Ioi, card_cons, Nat.add_sub_cancel_right] end OrderTop @@ -647,7 +629,7 @@ theorem not_mem_Iio_self {b : α} : b ∉ Iio b := fun h => lt_irrefl _ (mem_Iio theorem Iic_eq_cons_Iio (b : α) : Iic b = (Iio b).cons b not_mem_Iio_self := by classical rw [cons_eq_insert, Iio_insert] -theorem card_Iio_eq_card_Iic_sub_one (a : α) : (Iio a).card = (Iic a).card - 1 := by +theorem card_Iio_eq_card_Iic_sub_one (a : α) : #(Iio a) = #(Iic a) - 1 := by rw [Iic_eq_cons_Iio, card_cons, Nat.add_sub_cancel_right] end OrderBot @@ -715,25 +697,25 @@ theorem Ico_inter_Ico {a b c d : α} : Ico a b ∩ Ico c d = Ico (max a c) (min Set.Ico_inter_Ico] @[simp] -theorem Ico_filter_lt (a b c : α) : ((Ico a b).filter fun x => x < c) = Ico a (min b c) := by +theorem Ico_filter_lt (a b c : α) : {x ∈ Ico a b | x < c} = Ico a (min b c) := by cases le_total b c with | inl h => rw [Ico_filter_lt_of_right_le h, min_eq_left h] | inr h => rw [Ico_filter_lt_of_le_right h, min_eq_right h] @[simp] -theorem Ico_filter_le (a b c : α) : ((Ico a b).filter fun x => c ≤ x) = Ico (max a c) b := by +theorem Ico_filter_le (a b c : α) : {x ∈ Ico a b | c ≤ x} = Ico (max a c) b := by cases le_total a c with | inl h => rw [Ico_filter_le_of_left_le h, max_eq_right h] | inr h => rw [Ico_filter_le_of_le_left h, max_eq_left h] @[simp] -theorem Ioo_filter_lt (a b c : α) : (Ioo a b).filter (· < c) = Ioo a (min b c) := by +theorem Ioo_filter_lt (a b c : α) : {x ∈ Ioo a b | x < c} = Ioo a (min b c) := by ext simp [and_assoc] @[simp] theorem Iio_filter_lt {α} [LinearOrder α] [LocallyFiniteOrderBot α] (a b : α) : - (Iio a).filter (· < b) = Iio (min a b) := by + {x ∈ Iio a | x < b} = Iio (min a b) := by ext simp [and_assoc] @@ -933,7 +915,7 @@ lemma transGen_wcovBy_of_le [Preorder α] [LocallyFiniteOrder α] {x y : α} (hx TransGen (· ⩿ ·) x y := by -- We proceed by well-founded induction on the cardinality of `Icc x y`. -- It's impossible for the cardinality to be zero since `x ≤ y` - have : (Ico x y).card < (Icc x y).card := card_lt_card <| + have : #(Ico x y) < #(Icc x y) := card_lt_card <| ⟨Ico_subset_Icc_self, not_subset.mpr ⟨y, ⟨right_mem_Icc.mpr hxy, right_not_mem_Ico⟩⟩⟩ by_cases hxy' : y ≤ x -- If `y ≤ x`, then `x ⩿ y` @@ -943,16 +925,15 @@ lemma transGen_wcovBy_of_le [Preorder α] [LocallyFiniteOrder α] {x y : α} (hx induction hypothesis to show that `Relation.TransGen (· ⩿ ·) x z`. -/ · have h_non : (Ico x y).Nonempty := ⟨x, mem_Ico.mpr ⟨le_rfl, lt_of_le_not_le hxy hxy'⟩⟩ obtain ⟨z, z_mem, hz⟩ := (Ico x y).exists_maximal h_non - have z_card : (Icc x z).card <(Icc x y).card := calc - (Icc x z).card ≤ (Ico x y).card := - card_le_card <| Icc_subset_Ico_right (mem_Ico.mp z_mem).2 - _ < (Icc x y).card := this + have z_card := calc + #(Icc x z) ≤ #(Ico x y) := card_le_card <| Icc_subset_Ico_right (mem_Ico.mp z_mem).2 + _ < #(Icc x y) := this have h₁ := transGen_wcovBy_of_le (mem_Ico.mp z_mem).1 have h₂ : z ⩿ y := by refine ⟨(mem_Ico.mp z_mem).2.le, fun c hzc hcy ↦ hz c ?_ hzc⟩ exact mem_Ico.mpr <| ⟨(mem_Ico.mp z_mem).1.trans hzc.le, hcy⟩ exact .tail h₁ h₂ -termination_by (Icc x y).card +termination_by #(Icc x y) /-- In a locally finite preorder, `≤` is the transitive closure of `⩿`. -/ lemma le_iff_transGen_wcovBy [Preorder α] [LocallyFiniteOrder α] {x y : α} : @@ -976,7 +957,7 @@ lemma transGen_covBy_of_lt [Preorder α] [LocallyFiniteOrder α] {x y : α} (hxy -- `Ico x y` is a nonempty finset and so contains a maximal element `z` and -- `Ico x z` has cardinality strictly less than the cardinality of `Ico x y` obtain ⟨z, z_mem, hz⟩ := (Ico x y).exists_maximal h_non - have z_card : (Ico x z).card < (Ico x y).card := card_lt_card <| ssubset_iff_of_subset + have z_card : #(Ico x z) < #(Ico x y) := card_lt_card <| ssubset_iff_of_subset (Ico_subset_Ico le_rfl (mem_Ico.mp z_mem).2.le) |>.mpr ⟨z, z_mem, right_not_mem_Ico⟩ /- Since `z` is maximal in `Ico x y`, `z ⋖ y`. -/ have hzy : z ⋖ y := by @@ -990,7 +971,7 @@ lemma transGen_covBy_of_lt [Preorder α] [LocallyFiniteOrder α] {x y : α} (hxy `x ≤ z`), and since `z ⋖ y` we conclude that `x ⋖ y` , then `Relation.TransGen.single`. -/ · simp only [lt_iff_le_not_le, not_and, not_not] at hxz exact .single (hzy.of_le_of_lt (hxz (mem_Ico.mp z_mem).1) hxy) -termination_by (Ico x y).card +termination_by #(Ico x y) /-- In a locally finite preorder, `<` is the transitive closure of `⋖`. -/ lemma lt_iff_transGen_covBy [Preorder α] [LocallyFiniteOrder α] {x y : α} : diff --git a/Mathlib/Order/Interval/Finset/Box.lean b/Mathlib/Order/Interval/Finset/Box.lean index ae26532a1a735..902ac28210bbd 100644 --- a/Mathlib/Order/Interval/Finset/Box.lean +++ b/Mathlib/Order/Interval/Finset/Box.lean @@ -65,9 +65,9 @@ variable {α β : Type*} [OrderedRing α] [OrderedRing β] [LocallyFiniteOrder [DecidableEq α] [DecidableEq β] [@DecidableRel (α × β) (· ≤ ·)] @[simp] lemma card_box_succ (n : ℕ) : - (box (n + 1) : Finset (α × β)).card = - (Icc (-n.succ : α) n.succ).card * (Icc (-n.succ : β) n.succ).card - - (Icc (-n : α) n).card * (Icc (-n : β) n).card := by + #(box (n + 1) : Finset (α × β)) = + #(Icc (-n.succ : α) n.succ) * #(Icc (-n.succ : β) n.succ) - + #(Icc (-n : α) n) * #(Icc (-n : β) n) := by rw [box_succ_eq_sdiff, card_sdiff (Icc_neg_mono n.le_succ), Finset.card_Icc_prod, Finset.card_Icc_prod] rfl @@ -81,7 +81,7 @@ variable {n : ℕ} {x : ℤ × ℤ} attribute [norm_cast] toNat_ofNat -lemma card_box : ∀ {n}, n ≠ 0 → (box n : Finset (ℤ × ℤ)).card = 8 * n +lemma card_box : ∀ {n}, n ≠ 0 → #(box n : Finset (ℤ × ℤ)) = 8 * n | n + 1, _ => by simp_rw [Prod.card_box_succ, card_Icc, sub_neg_eq_add] norm_cast diff --git a/Mathlib/Order/Interval/Finset/Defs.lean b/Mathlib/Order/Interval/Finset/Defs.lean index 4957140ea5893..625bd3be33b31 100644 --- a/Mathlib/Order/Interval/Finset/Defs.lean +++ b/Mathlib/Order/Interval/Finset/Defs.lean @@ -70,7 +70,7 @@ Provide the `LocallyFiniteOrder` instance for `α →₀ β` where `β` is local From `LinearOrder α`, `NoMaxOrder α`, `LocallyFiniteOrder α`, we can also define an order isomorphism `α ≃ ℕ` or `α ≃ ℤ`, depending on whether we have `OrderBot α` or -`NoMinOrder α` and `Nonempty α`. When `OrderBot α`, we can match `a : α` to `(Iio a).card`. +`NoMinOrder α` and `Nonempty α`. When `OrderBot α`, we can match `a : α` to `#(Iio a)`. We can provide `SuccOrder α` from `LinearOrder α` and `LocallyFiniteOrder α` using @@ -146,79 +146,76 @@ the ends. As opposed to `LocallyFiniteOrder.ofIcc`, this one requires `Decidable only `Preorder`. -/ def LocallyFiniteOrder.ofIcc' (α : Type*) [Preorder α] [DecidableRel ((· ≤ ·) : α → α → Prop)] (finsetIcc : α → α → Finset α) (mem_Icc : ∀ a b x, x ∈ finsetIcc a b ↔ a ≤ x ∧ x ≤ b) : - LocallyFiniteOrder α := - { finsetIcc - finsetIco := fun a b => (finsetIcc a b).filter fun x => ¬b ≤ x - finsetIoc := fun a b => (finsetIcc a b).filter fun x => ¬x ≤ a - finsetIoo := fun a b => (finsetIcc a b).filter fun x => ¬x ≤ a ∧ ¬b ≤ x - finset_mem_Icc := mem_Icc - finset_mem_Ico := fun a b x => by rw [Finset.mem_filter, mem_Icc, and_assoc, lt_iff_le_not_le] - finset_mem_Ioc := fun a b x => by - rw [Finset.mem_filter, mem_Icc, and_right_comm, lt_iff_le_not_le] - finset_mem_Ioo := fun a b x => by - rw [Finset.mem_filter, mem_Icc, and_and_and_comm, lt_iff_le_not_le, lt_iff_le_not_le] } + LocallyFiniteOrder α where + finsetIcc := finsetIcc + finsetIco a b := {x ∈ finsetIcc a b | ¬b ≤ x} + finsetIoc a b := {x ∈ finsetIcc a b | ¬x ≤ a} + finsetIoo a b := {x ∈ finsetIcc a b | ¬x ≤ a ∧ ¬b ≤ x} + finset_mem_Icc := mem_Icc + finset_mem_Ico a b x := by rw [Finset.mem_filter, mem_Icc, and_assoc, lt_iff_le_not_le] + finset_mem_Ioc a b x := by rw [Finset.mem_filter, mem_Icc, and_right_comm, lt_iff_le_not_le] + finset_mem_Ioo a b x := by + rw [Finset.mem_filter, mem_Icc, and_and_and_comm, lt_iff_le_not_le, lt_iff_le_not_le] /-- A constructor from a definition of `Finset.Icc` alone, the other ones being derived by removing the ends. As opposed to `LocallyFiniteOrder.ofIcc'`, this one requires `PartialOrder` but only `DecidableEq`. -/ def LocallyFiniteOrder.ofIcc (α : Type*) [PartialOrder α] [DecidableEq α] (finsetIcc : α → α → Finset α) (mem_Icc : ∀ a b x, x ∈ finsetIcc a b ↔ a ≤ x ∧ x ≤ b) : - LocallyFiniteOrder α := - { finsetIcc - finsetIco := fun a b => (finsetIcc a b).filter fun x => x ≠ b - finsetIoc := fun a b => (finsetIcc a b).filter fun x => a ≠ x - finsetIoo := fun a b => (finsetIcc a b).filter fun x => a ≠ x ∧ x ≠ b - finset_mem_Icc := mem_Icc - finset_mem_Ico := fun a b x => by rw [Finset.mem_filter, mem_Icc, and_assoc, lt_iff_le_and_ne] - finset_mem_Ioc := fun a b x => by - rw [Finset.mem_filter, mem_Icc, and_right_comm, lt_iff_le_and_ne] - finset_mem_Ioo := fun a b x => by - rw [Finset.mem_filter, mem_Icc, and_and_and_comm, lt_iff_le_and_ne, lt_iff_le_and_ne] } + LocallyFiniteOrder α where + finsetIcc := finsetIcc + finsetIco a b := {x ∈ finsetIcc a b | x ≠ b} + finsetIoc a b := {x ∈ finsetIcc a b | a ≠ x} + finsetIoo a b := {x ∈ finsetIcc a b | a ≠ x ∧ x ≠ b} + finset_mem_Icc := mem_Icc + finset_mem_Ico a b x := by rw [Finset.mem_filter, mem_Icc, and_assoc, lt_iff_le_and_ne] + finset_mem_Ioc a b x := by rw [Finset.mem_filter, mem_Icc, and_right_comm, lt_iff_le_and_ne] + finset_mem_Ioo a b x := by + rw [Finset.mem_filter, mem_Icc, and_and_and_comm, lt_iff_le_and_ne, lt_iff_le_and_ne] /-- A constructor from a definition of `Finset.Ici` alone, the other ones being derived by removing the ends. As opposed to `LocallyFiniteOrderTop.ofIci`, this one requires `DecidableRel (· ≤ ·)` but only `Preorder`. -/ def LocallyFiniteOrderTop.ofIci' (α : Type*) [Preorder α] [DecidableRel ((· ≤ ·) : α → α → Prop)] (finsetIci : α → Finset α) (mem_Ici : ∀ a x, x ∈ finsetIci a ↔ a ≤ x) : - LocallyFiniteOrderTop α := - { finsetIci - finsetIoi := fun a => (finsetIci a).filter fun x => ¬x ≤ a - finset_mem_Ici := mem_Ici - finset_mem_Ioi := fun a x => by rw [mem_filter, mem_Ici, lt_iff_le_not_le] } + LocallyFiniteOrderTop α where + finsetIci := finsetIci + finsetIoi a := {x ∈ finsetIci a | ¬x ≤ a} + finset_mem_Ici := mem_Ici + finset_mem_Ioi a x := by rw [mem_filter, mem_Ici, lt_iff_le_not_le] /-- A constructor from a definition of `Finset.Ici` alone, the other ones being derived by removing the ends. As opposed to `LocallyFiniteOrderTop.ofIci'`, this one requires `PartialOrder` but only `DecidableEq`. -/ def LocallyFiniteOrderTop.ofIci (α : Type*) [PartialOrder α] [DecidableEq α] (finsetIci : α → Finset α) (mem_Ici : ∀ a x, x ∈ finsetIci a ↔ a ≤ x) : - LocallyFiniteOrderTop α := - { finsetIci - finsetIoi := fun a => (finsetIci a).filter fun x => a ≠ x - finset_mem_Ici := mem_Ici - finset_mem_Ioi := fun a x => by rw [mem_filter, mem_Ici, lt_iff_le_and_ne] } + LocallyFiniteOrderTop α where + finsetIci := finsetIci + finsetIoi a := {x ∈ finsetIci a | a ≠ x} + finset_mem_Ici := mem_Ici + finset_mem_Ioi a x := by rw [mem_filter, mem_Ici, lt_iff_le_and_ne] /-- A constructor from a definition of `Finset.Iic` alone, the other ones being derived by removing the ends. As opposed to `LocallyFiniteOrderBot.ofIic`, this one requires `DecidableRel (· ≤ ·)` but only `Preorder`. -/ def LocallyFiniteOrderBot.ofIic' (α : Type*) [Preorder α] [DecidableRel ((· ≤ ·) : α → α → Prop)] (finsetIic : α → Finset α) (mem_Iic : ∀ a x, x ∈ finsetIic a ↔ x ≤ a) : - LocallyFiniteOrderBot α := - { finsetIic - finsetIio := fun a => (finsetIic a).filter fun x => ¬a ≤ x - finset_mem_Iic := mem_Iic - finset_mem_Iio := fun a x => by rw [mem_filter, mem_Iic, lt_iff_le_not_le] } + LocallyFiniteOrderBot α where + finsetIic := finsetIic + finsetIio a := {x ∈ finsetIic a | ¬a ≤ x} + finset_mem_Iic := mem_Iic + finset_mem_Iio a x := by rw [mem_filter, mem_Iic, lt_iff_le_not_le] /-- A constructor from a definition of `Finset.Iic` alone, the other ones being derived by removing the ends. As opposed to `LocallyFiniteOrderBot.ofIic'`, this one requires `PartialOrder` but only `DecidableEq`. -/ def LocallyFiniteOrderBot.ofIic (α : Type*) [PartialOrder α] [DecidableEq α] (finsetIic : α → Finset α) (mem_Iic : ∀ a x, x ∈ finsetIic a ↔ x ≤ a) : - LocallyFiniteOrderBot α := - { finsetIic - finsetIio := fun a => (finsetIic a).filter fun x => x ≠ a - finset_mem_Iic := mem_Iic - finset_mem_Iio := fun a x => by rw [mem_filter, mem_Iic, lt_iff_le_and_ne] } --- Note: this was in the wrong namespace in Mathlib 3. + LocallyFiniteOrderBot α where + finsetIic := finsetIic + finsetIio a := {x ∈ finsetIic a | x ≠ a} + finset_mem_Iic := mem_Iic + finset_mem_Iio a x := by rw [mem_filter, mem_Iic, lt_iff_le_and_ne] variable {α β : Type*} @@ -786,8 +783,8 @@ lemma Finset.Icc_prod_def (x y : α × β) : Icc x y = Icc x.1 y.1 ×ˢ Icc x.2 lemma Finset.Icc_product_Icc (a₁ a₂ : α) (b₁ b₂ : β) : Icc a₁ a₂ ×ˢ Icc b₁ b₂ = Icc (a₁, b₁) (a₂, b₂) := rfl -lemma Finset.card_Icc_prod (x y : α × β) : - (Icc x y).card = (Icc x.1 y.1).card * (Icc x.2 y.2).card := card_product _ _ +lemma Finset.card_Icc_prod (x y : α × β) : #(Icc x y) = #(Icc x.1 y.1) * #(Icc x.2 y.2) := + card_product .. end LocallyFiniteOrder @@ -800,7 +797,7 @@ instance Prod.instLocallyFiniteOrderTop : LocallyFiniteOrderTop (α × β) := lemma Finset.Ici_prod_def (x : α × β) : Ici x = Ici x.1 ×ˢ Ici x.2 := rfl lemma Finset.Ici_product_Ici (a : α) (b : β) : Ici a ×ˢ Ici b = Ici (a, b) := rfl -lemma Finset.card_Ici_prod (x : α × β) : (Ici x).card = (Ici x.1).card * (Ici x.2).card := +lemma Finset.card_Ici_prod (x : α × β) : #(Ici x) = #(Ici x.1) * #(Ici x.2) := card_product _ _ end LocallyFiniteOrderTop @@ -814,8 +811,7 @@ instance Prod.instLocallyFiniteOrderBot : LocallyFiniteOrderBot (α × β) := lemma Finset.Iic_prod_def (x : α × β) : Iic x = Iic x.1 ×ˢ Iic x.2 := rfl lemma Finset.Iic_product_Iic (a : α) (b : β) : Iic a ×ˢ Iic b = Iic (a, b) := rfl -lemma Finset.card_Iic_prod (x : α × β) : (Iic x).card = (Iic x.1).card * (Iic x.2).card := - card_product _ _ +lemma Finset.card_Iic_prod (x : α × β) : #(Iic x) = #(Iic x.1) * #(Iic x.2) := card_product .. end LocallyFiniteOrderBot end Preorder @@ -829,8 +825,8 @@ lemma Finset.uIcc_prod_def (x y : α × β) : uIcc x y = uIcc x.1 y.1 ×ˢ uIcc lemma Finset.uIcc_product_uIcc (a₁ a₂ : α) (b₁ b₂ : β) : uIcc a₁ a₂ ×ˢ uIcc b₁ b₂ = uIcc (a₁, b₁) (a₂, b₂) := rfl -lemma Finset.card_uIcc_prod (x y : α × β) : - (uIcc x y).card = (uIcc x.1 y.1).card * (uIcc x.2 y.2).card := card_product _ _ +lemma Finset.card_uIcc_prod (x y : α × β) : #(uIcc x y) = #(uIcc x.1 y.1) * #(uIcc x.2 y.2) := + card_product .. end Lattice diff --git a/Mathlib/Order/Interval/Finset/Fin.lean b/Mathlib/Order/Interval/Finset/Fin.lean index f0a6f5f7b6aef..24bf50afbdc49 100644 --- a/Mathlib/Order/Interval/Finset/Fin.lean +++ b/Mathlib/Order/Interval/Finset/Fin.lean @@ -85,23 +85,19 @@ theorem map_subtype_embedding_uIcc : (uIcc a b).map valEmbedding = uIcc ↑a ↑ map_valEmbedding_Icc _ _ @[simp] -theorem card_Icc : (Icc a b).card = b + 1 - a := by - rw [← Nat.card_Icc, ← map_valEmbedding_Icc, card_map] +lemma card_Icc : #(Icc a b) = b + 1 - a := by rw [← Nat.card_Icc, ← map_valEmbedding_Icc, card_map] @[simp] -theorem card_Ico : (Ico a b).card = b - a := by - rw [← Nat.card_Ico, ← map_valEmbedding_Ico, card_map] +lemma card_Ico : #(Ico a b) = b - a := by rw [← Nat.card_Ico, ← map_valEmbedding_Ico, card_map] @[simp] -theorem card_Ioc : (Ioc a b).card = b - a := by - rw [← Nat.card_Ioc, ← map_valEmbedding_Ioc, card_map] +lemma card_Ioc : #(Ioc a b) = b - a := by rw [← Nat.card_Ioc, ← map_valEmbedding_Ioc, card_map] @[simp] -theorem card_Ioo : (Ioo a b).card = b - a - 1 := by - rw [← Nat.card_Ioo, ← map_valEmbedding_Ioo, card_map] +lemma card_Ioo : #(Ioo a b) = b - a - 1 := by rw [← Nat.card_Ioo, ← map_valEmbedding_Ioo, card_map] @[simp] -theorem card_uIcc : (uIcc a b).card = (b - a : ℤ).natAbs + 1 := by +theorem card_uIcc : #(uIcc a b) = (b - a : ℤ).natAbs + 1 := by rw [← Nat.card_uIcc, ← map_subtype_embedding_uIcc, card_map] -- Porting note (#10618): simp can prove this @@ -172,23 +168,20 @@ theorem map_valEmbedding_Iio : (Iio b).map Fin.valEmbedding = Iio ↑b := by simp [Iio_eq_finset_subtype, Finset.fin, Finset.map_map] @[simp] -theorem card_Ici : (Ici a).card = n - a := by +theorem card_Ici : #(Ici a) = n - a := by cases n with | zero => exact Fin.elim0 a | succ => rw [← card_map, map_valEmbedding_Ici, Nat.card_Icc, Nat.add_one_sub_one] @[simp] -theorem card_Ioi : (Ioi a).card = n - 1 - a := by - rw [← card_map, map_valEmbedding_Ioi, Nat.card_Ioc] +theorem card_Ioi : #(Ioi a) = n - 1 - a := by rw [← card_map, map_valEmbedding_Ioi, Nat.card_Ioc] @[simp] -theorem card_Iic : (Iic b).card = b + 1 := by - rw [← Nat.card_Iic b, ← map_valEmbedding_Iic, card_map] +theorem card_Iic : #(Iic b) = b + 1 := by rw [← Nat.card_Iic b, ← map_valEmbedding_Iic, card_map] @[simp] -theorem card_Iio : (Iio b).card = b := by - rw [← Nat.card_Iio b, ← map_valEmbedding_Iio, card_map] +theorem card_Iio : #(Iio b) = b := by rw [← Nat.card_Iio b, ← map_valEmbedding_Iio, card_map] -- Porting note (#10618): simp can prove this -- @[simp] diff --git a/Mathlib/Order/Interval/Finset/Nat.lean b/Mathlib/Order/Interval/Finset/Nat.lean index 8be56194d6481..6b580e5787485 100644 --- a/Mathlib/Order/Interval/Finset/Nat.lean +++ b/Mathlib/Order/Interval/Finset/Nat.lean @@ -65,31 +65,20 @@ lemma range_eq_Icc_zero_sub_one (n : ℕ) (hn : n ≠ 0) : range n = Icc 0 (n - theorem _root_.Finset.range_eq_Ico : range = Ico 0 := Ico_zero_eq_range.symm -@[simp] -theorem card_Icc : (Icc a b).card = b + 1 - a := - List.length_range' _ _ _ - -@[simp] -theorem card_Ico : (Ico a b).card = b - a := - List.length_range' _ _ _ - -@[simp] -theorem card_Ioc : (Ioc a b).card = b - a := - List.length_range' _ _ _ - -@[simp] -theorem card_Ioo : (Ioo a b).card = b - a - 1 := - List.length_range' _ _ _ +@[simp] lemma card_Icc : #(Icc a b) = b + 1 - a := List.length_range' .. +@[simp] lemma card_Ico : #(Ico a b) = b - a := List.length_range' .. +@[simp] lemma card_Ioc : #(Ioc a b) = b - a := List.length_range' .. +@[simp] lemma card_Ioo : #(Ioo a b) = b - a - 1 := List.length_range' .. @[simp] -theorem card_uIcc : (uIcc a b).card = (b - a : ℤ).natAbs + 1 := +theorem card_uIcc : #(uIcc a b) = (b - a : ℤ).natAbs + 1 := (card_Icc _ _).trans <| by rw [← Int.natCast_inj, sup_eq_max, inf_eq_min, Int.ofNat_sub] <;> omega @[simp] -lemma card_Iic : (Iic b).card = b + 1 := by rw [Iic_eq_Icc, card_Icc, Nat.bot_eq_zero, Nat.sub_zero] +lemma card_Iic : #(Iic b) = b + 1 := by rw [Iic_eq_Icc, card_Icc, Nat.bot_eq_zero, Nat.sub_zero] @[simp] -theorem card_Iio : (Iio b).card = b := by rw [Iio_eq_Ico, card_Ico, Nat.bot_eq_zero, Nat.sub_zero] +theorem card_Iio : #(Iio b) = b := by rw [Iio_eq_Ico, card_Ico, Nat.bot_eq_zero, Nat.sub_zero] -- Porting note (#10618): simp can prove this -- @[simp] diff --git a/Mathlib/Order/Partition/Equipartition.lean b/Mathlib/Order/Partition/Equipartition.lean index 3969b84ca5eb0..4cd1f1579f339 100644 --- a/Mathlib/Order/Partition/Equipartition.lean +++ b/Mathlib/Order/Partition/Equipartition.lean @@ -35,44 +35,41 @@ def IsEquipartition : Prop := theorem isEquipartition_iff_card_parts_eq_average : P.IsEquipartition ↔ - ∀ a : Finset α, - a ∈ P.parts → a.card = s.card / P.parts.card ∨ a.card = s.card / P.parts.card + 1 := by + ∀ a : Finset α, a ∈ P.parts → #a = #s / #P.parts ∨ #a = #s / #P.parts + 1 := by simp_rw [IsEquipartition, Finset.equitableOn_iff, P.sum_card_parts] variable {P} lemma not_isEquipartition : - ¬P.IsEquipartition ↔ ∃ a ∈ P.parts, ∃ b ∈ P.parts, b.card + 1 < a.card := - Set.not_equitableOn + ¬P.IsEquipartition ↔ ∃ a ∈ P.parts, ∃ b ∈ P.parts, #b + 1 < #a := Set.not_equitableOn theorem _root_.Set.Subsingleton.isEquipartition (h : (P.parts : Set (Finset α)).Subsingleton) : P.IsEquipartition := Set.Subsingleton.equitableOn h _ theorem IsEquipartition.card_parts_eq_average (hP : P.IsEquipartition) (ht : t ∈ P.parts) : - t.card = s.card / P.parts.card ∨ t.card = s.card / P.parts.card + 1 := + #t = #s / #P.parts ∨ #t = #s / #P.parts + 1 := P.isEquipartition_iff_card_parts_eq_average.1 hP _ ht theorem IsEquipartition.card_part_eq_average_iff (hP : P.IsEquipartition) (ht : t ∈ P.parts) : - t.card = s.card / P.parts.card ↔ t.card ≠ s.card / P.parts.card + 1 := by + #t = #s / #P.parts ↔ #t ≠ #s / #P.parts + 1 := by have a := hP.card_parts_eq_average ht - have b : ¬(t.card = s.card / P.parts.card ∧ t.card = s.card / P.parts.card + 1) := by + have b : ¬(#t = #s / #P.parts ∧ #t = #s / #P.parts + 1) := by by_contra h; exact absurd (h.1 ▸ h.2) (lt_add_one _).ne tauto theorem IsEquipartition.average_le_card_part (hP : P.IsEquipartition) (ht : t ∈ P.parts) : - s.card / P.parts.card ≤ t.card := by + #s / #P.parts ≤ #t := by rw [← P.sum_card_parts] exact Finset.EquitableOn.le hP ht theorem IsEquipartition.card_part_le_average_add_one (hP : P.IsEquipartition) (ht : t ∈ P.parts) : - t.card ≤ s.card / P.parts.card + 1 := by + #t ≤ #s / #P.parts + 1 := by rw [← P.sum_card_parts] exact Finset.EquitableOn.le_add_one hP ht theorem IsEquipartition.filter_ne_average_add_one_eq_average (hP : P.IsEquipartition) : - P.parts.filter (fun p ↦ ¬p.card = s.card / P.parts.card + 1) = - P.parts.filter (fun p ↦ p.card = s.card / P.parts.card) := by + {p ∈ P.parts | ¬#p = #s / #P.parts + 1} = {p ∈ P.parts | #p = #s / #P.parts} := by ext p simp only [mem_filter, and_congr_right_iff] exact fun hp ↦ (hP.card_part_eq_average_iff hp).symm @@ -80,75 +77,70 @@ theorem IsEquipartition.filter_ne_average_add_one_eq_average (hP : P.IsEquiparti /-- An equipartition of a finset with `n` elements into `k` parts has `n % k` parts of size `n / k + 1`. -/ theorem IsEquipartition.card_large_parts_eq_mod (hP : P.IsEquipartition) : - (P.parts.filter fun p ↦ p.card = s.card / P.parts.card + 1).card = s.card % P.parts.card := by + #{p ∈ P.parts | #p = #s / #P.parts + 1} = #s % #P.parts := by have z := P.sum_card_parts - rw [← sum_filter_add_sum_filter_not (s := P.parts) - (p := fun x ↦ x.card = s.card / P.parts.card + 1), - hP.filter_ne_average_add_one_eq_average, - sum_const_nat (m := s.card / P.parts.card + 1) (by simp), - sum_const_nat (m := s.card / P.parts.card) (by simp), - ← hP.filter_ne_average_add_one_eq_average, - mul_add, add_comm, ← add_assoc, ← add_mul, mul_one, add_comm (Finset.card _), + rw [← sum_filter_add_sum_filter_not (s := P.parts) (p := fun x ↦ #x = #s / #P.parts + 1), + hP.filter_ne_average_add_one_eq_average, sum_const_nat (m := #s / #P.parts + 1) (by simp), + sum_const_nat (m := #s / #P.parts) (by simp), ← hP.filter_ne_average_add_one_eq_average, + mul_add, add_comm, ← add_assoc, ← add_mul, mul_one, add_comm #_, filter_card_add_filter_neg_card_eq_card, add_comm] at z rw [← add_left_inj, Nat.mod_add_div, z] /-- An equipartition of a finset with `n` elements into `k` parts has `n - n % k` parts of size `n / k`. -/ theorem IsEquipartition.card_small_parts_eq_mod (hP : P.IsEquipartition) : - (P.parts.filter fun p ↦ p.card = s.card / P.parts.card).card = - P.parts.card - s.card % P.parts.card := by + #{p ∈ P.parts | #p = #s / #P.parts} = #P.parts - #s % #P.parts := by conv_rhs => arg 1 - rw [← filter_card_add_filter_neg_card_eq_card (p := fun p ↦ p.card = s.card / P.parts.card + 1)] + rw [← filter_card_add_filter_neg_card_eq_card (p := fun p ↦ #p = #s / #P.parts + 1)] rw [hP.card_large_parts_eq_mod, add_tsub_cancel_left, hP.filter_ne_average_add_one_eq_average] /-- There exists an enumeration of an equipartition's parts where larger parts map to smaller numbers and vice versa. -/ theorem IsEquipartition.exists_partsEquiv (hP : P.IsEquipartition) : - ∃ f : P.parts ≃ Fin P.parts.card, - ∀ t, t.1.card = s.card / P.parts.card + 1 ↔ f t < s.card % P.parts.card := by - let el := (P.parts.filter fun p ↦ p.card = s.card / P.parts.card + 1).equivFin - let es := (P.parts.filter fun p ↦ p.card = s.card / P.parts.card).equivFin + ∃ f : P.parts ≃ Fin #P.parts, ∀ t, #t.1 = #s / #P.parts + 1 ↔ f t < #s % #P.parts := by + let el := {p ∈ P.parts | #p = #s / #P.parts + 1}.equivFin + let es := {p ∈ P.parts | #p = #s / #P.parts}.equivFin simp_rw [mem_filter, hP.card_large_parts_eq_mod] at el simp_rw [mem_filter, hP.card_small_parts_eq_mod] at es - let sneg : { x // x ∈ P.parts ∧ ¬x.card = s.card / P.parts.card + 1 } ≃ - { x // x ∈ P.parts ∧ x.card = s.card / P.parts.card } := by + let sneg : + {x // x ∈ P.parts ∧ ¬#x = #s / #P.parts + 1} ≃ {x // x ∈ P.parts ∧ #x = #s / #P.parts} := by apply (Equiv.refl _).subtypeEquiv simp only [Equiv.refl_apply, and_congr_right_iff] exact fun _ ha ↦ by rw [hP.card_part_eq_average_iff ha, ne_eq] - replace el : { x : P.parts // x.1.card = s.card / P.parts.card + 1 } ≃ - Fin (s.card % P.parts.card) := (Equiv.Set.sep ..).symm.trans el - replace es : { x : P.parts // ¬x.1.card = s.card / P.parts.card + 1 } ≃ - Fin (P.parts.card - s.card % P.parts.card) := (Equiv.Set.sep ..).symm.trans (sneg.trans es) + replace el : { x : P.parts // #x.1 = #s / #P.parts + 1 } ≃ + Fin (#s % #P.parts) := (Equiv.Set.sep ..).symm.trans el + replace es : { x : P.parts // ¬#x.1 = #s / #P.parts + 1 } ≃ + Fin (#P.parts - #s % #P.parts) := (Equiv.Set.sep ..).symm.trans (sneg.trans es) let f := (Equiv.sumCompl _).symm.trans ((el.sumCongr es).trans finSumFinEquiv) use f.trans (finCongr (Nat.add_sub_of_le P.card_mod_card_parts_le)) intro ⟨p, _⟩ simp_rw [f, Equiv.trans_apply, Equiv.sumCongr_apply, finCongr_apply, Fin.coe_cast] - by_cases hc : p.card = s.card / P.parts.card + 1 <;> simp [hc] + by_cases hc : #p = #s / #P.parts + 1 <;> simp [hc] /-- Given a finset equipartitioned into `k` parts, its elements can be enumerated such that elements in the same part have congruent indices modulo `k`. -/ -theorem IsEquipartition.exists_partPreservingEquiv (hP : P.IsEquipartition) : ∃ f : s ≃ Fin s.card, - ∀ a b : s, P.part a = P.part b ↔ f a % P.parts.card = f b % P.parts.card := by +theorem IsEquipartition.exists_partPreservingEquiv (hP : P.IsEquipartition) : ∃ f : s ≃ Fin #s, + ∀ a b : s, P.part a = P.part b ↔ f a % #P.parts = f b % #P.parts := by obtain ⟨f, hf⟩ := P.exists_enumeration obtain ⟨g, hg⟩ := hP.exists_partsEquiv - let z := fun a ↦ P.parts.card * (f a).2 + g (f a).1 + let z := fun a ↦ #P.parts * (f a).2 + g (f a).1 have gl := fun a ↦ (g (f a).1).2 - have less : ∀ a, z a < s.card := fun a ↦ by + have less : ∀ a, z a < #s := fun a ↦ by rcases hP.card_parts_eq_average (f a).1.2 with (c | c) · calc - _ < P.parts.card * ((f a).2 + 1) := add_lt_add_left (gl a) _ - _ ≤ P.parts.card * (s.card / P.parts.card) := mul_le_mul_left' (c ▸ (f a).2.2) _ - _ ≤ P.parts.card * (s.card / P.parts.card) + s.card % P.parts.card := Nat.le_add_right .. + _ < #P.parts * ((f a).2 + 1) := add_lt_add_left (gl a) _ + _ ≤ #P.parts * (#s / #P.parts) := mul_le_mul_left' (c ▸ (f a).2.2) _ + _ ≤ #P.parts * (#s / #P.parts) + #s % #P.parts := Nat.le_add_right .. _ = _ := Nat.div_add_mod .. - · rw [← Nat.div_add_mod s.card P.parts.card] + · rw [← Nat.div_add_mod #s #P.parts] exact add_lt_add_of_le_of_lt (mul_le_mul_left' (by omega) _) ((hg (f a).1).mp c) - let z' : s → Fin s.card := fun a ↦ ⟨z a, less a⟩ + let z' : s → Fin #s := fun a ↦ ⟨z a, less a⟩ have bij : z'.Bijective := by refine (bijective_iff_injective_and_card z').mpr ⟨fun a b e ↦ ?_, by simp⟩ - simp_rw [z', z, Fin.mk.injEq, mul_comm P.parts.card] at e - haveI : NeZero P.parts.card := ⟨((Nat.zero_le _).trans_lt (gl a)).ne'⟩ - change P.parts.card.divModEquiv.symm (_, _) = P.parts.card.divModEquiv.symm (_, _) at e + simp_rw [z', z, Fin.mk.injEq, mul_comm #P.parts] at e + haveI : NeZero #P.parts := ⟨((Nat.zero_le _).trans_lt (gl a)).ne'⟩ + change (#P.parts).divModEquiv.symm (_, _) = (#P.parts).divModEquiv.symm (_, _) at e simp only [Equiv.apply_eq_iff_eq, Prod.mk.injEq] at e apply_fun f exact Sigma.ext e.2 <| (Fin.heq_ext_iff (by rw [e.2])).mpr e.1 diff --git a/Mathlib/Order/Partition/Finpartition.lean b/Mathlib/Order/Partition/Finpartition.lean index 0c1bbdebaec82..f92a6209a1310 100644 --- a/Mathlib/Order/Partition/Finpartition.lean +++ b/Mathlib/Order/Partition/Finpartition.lean @@ -338,7 +338,7 @@ theorem exists_le_of_le {a b : α} {P Q : Finpartition a} (h : P ≤ Q) (hb : b simp only [not_exists, not_and] at H exact H _ hc hcd -theorem card_mono {a : α} {P Q : Finpartition a} (h : P ≤ Q) : Q.parts.card ≤ P.parts.card := by +theorem card_mono {a : α} {P Q : Finpartition a} (h : P ≤ Q) : #Q.parts ≤ #P.parts := by classical have : ∀ b ∈ Q.parts, ∃ c ∈ P.parts, c ≤ b := fun b ↦ exists_le_of_le h choose f hP hf using this @@ -389,7 +389,7 @@ theorem mem_bind : b ∈ (P.bind Q).parts ↔ ∃ A hA, b ∈ (Q A hA).parts := exact ⟨⟨A, hA⟩, mem_attach _ ⟨A, hA⟩, h⟩ theorem card_bind (Q : ∀ i ∈ P.parts, Finpartition i) : - (P.bind Q).parts.card = ∑ A ∈ P.parts.attach, (Q _ A.2).parts.card := by + #(P.bind Q).parts = ∑ A ∈ P.parts.attach, #(Q _ A.2).parts := by apply card_biUnion rintro ⟨b, hb⟩ - ⟨c, hc⟩ - hbc rw [Finset.disjoint_left] @@ -414,7 +414,7 @@ def extend (P : Finpartition a) (hb : b ≠ ⊥) (hab : Disjoint a b) (hc : a not_bot_mem h := (mem_insert.1 h).elim hb.symm P.not_bot_mem theorem card_extend (P : Finpartition a) (b c : α) {hb : b ≠ ⊥} {hab : Disjoint a b} - {hc : a ⊔ b = c} : (P.extend hb hab hc).parts.card = P.parts.card + 1 := + {hc : a ⊔ b = c} : #(P.extend hb hab hc).parts = #P.parts + 1 := card_insert_of_not_mem fun h ↦ hb <| hab.symm.eq_bot_of_le <| P.le h end DistribLattice @@ -506,12 +506,12 @@ def equivSigmaParts : s ≃ Σ t : P.parts, t.1 where rw [P.part_eq_of_mem mp mf] · simp -lemma exists_enumeration : ∃ f : s ≃ Σ t : P.parts, Fin t.1.card, +lemma exists_enumeration : ∃ f : s ≃ Σ t : P.parts, Fin #t.1, ∀ a b : s, P.part a = P.part b ↔ (f a).1 = (f b).1 := by use P.equivSigmaParts.trans ((Equiv.refl _).sigmaCongr (fun t ↦ t.1.equivFin)) simp [equivSigmaParts, Equiv.sigmaCongr, Equiv.sigmaCongrLeft] -theorem sum_card_parts : ∑ i ∈ P.parts, i.card = s.card := by +theorem sum_card_parts : ∑ i ∈ P.parts, #i = #s := by convert congr_arg Finset.card P.biUnion_parts rw [card_biUnion P.supIndep.pairwiseDisjoint] rfl @@ -532,8 +532,7 @@ theorem parts_bot (s : Finset α) : (⊥ : Finpartition s).parts = s.map ⟨singleton, singleton_injective⟩ := rfl -theorem card_bot (s : Finset α) : (⊥ : Finpartition s).parts.card = s.card := - Finset.card_map _ +theorem card_bot (s : Finset α) : #(⊥ : Finpartition s).parts = #s := Finset.card_map _ theorem mem_bot_iff : t ∈ (⊥ : Finpartition s).parts ↔ ∃ a ∈ s, {a} = t := mem_map @@ -546,15 +545,15 @@ instance (s : Finset α) : OrderBot (Finpartition s) := obtain ⟨t, ht, hat⟩ := P.exists_mem ha exact ⟨t, ht, singleton_subset_iff.2 hat⟩ } -theorem card_parts_le_card : P.parts.card ≤ s.card := by +theorem card_parts_le_card : #P.parts ≤ #s := by rw [← card_bot s] exact card_mono bot_le -lemma card_mod_card_parts_le : s.card % P.parts.card ≤ P.parts.card := by - rcases P.parts.card.eq_zero_or_pos with h | h - · have h' := h - rw [Finset.card_eq_zero, parts_eq_empty_iff, bot_eq_empty, ← Finset.card_eq_zero] at h' - rw [h, h'] +lemma card_mod_card_parts_le : #s % #P.parts ≤ #P.parts := by + obtain h | h := (#P.parts).eq_zero_or_pos + · rw [h] + rw [Finset.card_eq_zero, parts_eq_empty_iff, bot_eq_empty, ← Finset.card_eq_zero] at h + rw [h] · exact (Nat.mod_lt _ h).le section Setoid @@ -564,7 +563,7 @@ variable [Fintype α] /-- A setoid over a finite type induces a finpartition of the type's elements, where the parts are the setoid's equivalence classes. -/ def ofSetoid (s : Setoid α) [DecidableRel s.r] : Finpartition (univ : Finset α) where - parts := univ.image fun a => univ.filter (s.r a) + parts := univ.image fun a ↦ ({b | s.r a b} : Finset α) supIndep := by simp only [mem_univ, forall_true_left, supIndep_iff_pairwiseDisjoint, Set.PairwiseDisjoint, Set.Pairwise, coe_image, coe_univ, Set.image_univ, Set.mem_range, ne_eq, @@ -605,7 +604,7 @@ section Atomise /-- Cuts `s` along the finsets in `F`: Two elements of `s` will be in the same part if they are in the same finsets of `F`. -/ def atomise (s : Finset α) (F : Finset (Finset α)) : Finpartition s := - ofErase (F.powerset.image fun Q ↦ s.filter fun i ↦ ∀ t ∈ F, t ∈ Q ↔ i ∈ t) + ofErase (F.powerset.image fun Q ↦ {i ∈ s | ∀ t ∈ F, t ∈ Q ↔ i ∈ t}) (Set.PairwiseDisjoint.supIndep fun x hx y hy h ↦ disjoint_left.mpr fun z hz1 hz2 ↦ h (by @@ -614,8 +613,7 @@ def atomise (s : Finset α) (F : Finset (Finset α)) : Finpartition s := obtain ⟨R, hR, rfl⟩ := hy suffices h' : Q = R by subst h' - exact of_eq_true (eq_self ( - filter (fun i ↦ ∀ (t : Finset α), t ∈ F → (t ∈ Q ↔ i ∈ t)) s)) + exact of_eq_true (eq_self {i ∈ s | ∀ t ∈ F, t ∈ Q ↔ i ∈ t}) rw [id, mem_filter] at hz1 hz2 rw [mem_powerset] at hQ hR ext i @@ -629,7 +627,7 @@ def atomise (s : Finset α) (F : Finset (Finset α)) : Finpartition s := exact s.filter_subset _ · rw [mem_sup] refine - ⟨s.filter fun i ↦ ∀ t, t ∈ F → ((t ∈ F.filter fun u ↦ a ∈ u) ↔ i ∈ t), + ⟨{i ∈ s | ∀ t ∈ F, t ∈ {u ∈ F | a ∈ u} ↔ i ∈ t}, mem_image_of_mem _ (mem_powerset.2 <| filter_subset _ _), mem_filter.2 ⟨ha, fun t ht ↦ ?_⟩⟩ rw [mem_filter] @@ -639,7 +637,7 @@ variable {F : Finset (Finset α)} theorem mem_atomise : t ∈ (atomise s F).parts ↔ - t.Nonempty ∧ ∃ Q ⊆ F, (s.filter fun i ↦ ∀ u ∈ F, u ∈ Q ↔ i ∈ u) = t := by + t.Nonempty ∧ ∃ Q ⊆ F, {i ∈ s | ∀ u ∈ F, u ∈ Q ↔ i ∈ u} = t := by simp only [atomise, ofErase, bot_eq_empty, mem_erase, mem_image, nonempty_iff_ne_empty, mem_singleton, and_comm, mem_powerset, exists_prop] @@ -648,11 +646,11 @@ theorem atomise_empty (hs : s.Nonempty) : (atomise s ∅).parts = {s} := by imp_true_iff, filter_True] exact erase_eq_of_not_mem (not_mem_singleton.2 hs.ne_empty.symm) -theorem card_atomise_le : (atomise s F).parts.card ≤ 2 ^ F.card := +theorem card_atomise_le : #(atomise s F).parts ≤ 2 ^ #F := (card_le_card <| erase_subset _ _).trans <| Finset.card_image_le.trans (card_powerset _).le theorem biUnion_filter_atomise (ht : t ∈ F) (hts : t ⊆ s) : - ((atomise s F).parts.filter fun u ↦ u ⊆ t ∧ u.Nonempty).biUnion id = t := by + {u ∈ (atomise s F).parts | u ⊆ t ∧ u.Nonempty}.biUnion id = t := by ext a refine mem_biUnion.trans ⟨fun ⟨u, hu, ha⟩ ↦ (mem_filter.1 hu).2.1 ha, fun ha ↦ ?_⟩ obtain ⟨u, hu, hau⟩ := (atomise s F).exists_mem (hts ha) @@ -662,10 +660,10 @@ theorem biUnion_filter_atomise (ht : t ∈ F) (hts : t ⊆ s) : rwa [← hb.2 _ ht, hau.2 _ ht] theorem card_filter_atomise_le_two_pow (ht : t ∈ F) : - ((atomise s F).parts.filter fun u ↦ u ⊆ t ∧ u.Nonempty).card ≤ 2 ^ (F.card - 1) := by + #{u ∈ (atomise s F).parts | u ⊆ t ∧ u.Nonempty} ≤ 2 ^ (#F - 1) := by suffices h : - ((atomise s F).parts.filter fun u ↦ u ⊆ t ∧ u.Nonempty) ⊆ - (F.erase t).powerset.image fun P ↦ s.filter fun i ↦ ∀ x ∈ F, x ∈ insert t P ↔ i ∈ x by + {u ∈ (atomise s F).parts | u ⊆ t ∧ u.Nonempty} ⊆ + (F.erase t).powerset.image fun P ↦ {i ∈ s | ∀ x ∈ F, x ∈ insert t P ↔ i ∈ x} by refine (card_le_card h).trans (card_image_le.trans ?_) rw [card_powerset, card_erase_of_mem ht] rw [subset_iff] diff --git a/Mathlib/Order/SupClosed.lean b/Mathlib/Order/SupClosed.lean index f1940a4250ed1..1e09bd963d671 100644 --- a/Mathlib/Order/SupClosed.lean +++ b/Mathlib/Order/SupClosed.lean @@ -306,7 +306,7 @@ lemma supClosure_min : s ⊆ t → SupClosed t → supClosure s ⊆ t := supClos protected lemma Set.Finite.supClosure (hs : s.Finite) : (supClosure s).Finite := by lift s to Finset α using hs classical - refine ((s.powerset.filter Finset.Nonempty).attach.image + refine ({t ∈ s.powerset | t.Nonempty}.attach.image fun t ↦ t.1.sup' (mem_filter.1 t.2).2 id).finite_toSet.subset ?_ rintro _ ⟨t, ht, hts, rfl⟩ simp only [id_eq, coe_image, mem_image, mem_coe, mem_attach, true_and, Subtype.exists, @@ -378,7 +378,7 @@ lemma infClosure_min : s ⊆ t → InfClosed t → infClosure s ⊆ t := infClos protected lemma Set.Finite.infClosure (hs : s.Finite) : (infClosure s).Finite := by lift s to Finset α using hs classical - refine ((s.powerset.filter Finset.Nonempty).attach.image + refine ({t ∈ s.powerset | t.Nonempty}.attach.image fun t ↦ t.1.inf' (mem_filter.1 t.2).2 id).finite_toSet.subset ?_ rintro _ ⟨t, ht, hts, rfl⟩ simp only [id_eq, coe_image, mem_image, mem_coe, mem_attach, true_and, Subtype.exists, diff --git a/Mathlib/Probability/StrongLaw.lean b/Mathlib/Probability/StrongLaw.lean index fae0c8332ee2f..4158b48564a82 100644 --- a/Mathlib/Probability/StrongLaw.lean +++ b/Mathlib/Probability/StrongLaw.lean @@ -420,8 +420,7 @@ theorem strong_law_aux1 {c : ℝ} (c_one : 1 < c) {ε : ℝ} (εpos : 0 < ε) : exact (hident j).aestronglyMeasurable_fst.memℒp_truncation · intro k _ l _ hkl exact (hindep hkl).comp (A k).measurable (A l).measurable - _ = ∑ j ∈ range (u (N - 1)), - (∑ i ∈ (range N).filter fun i => j < u i, ((u i : ℝ) ^ 2)⁻¹) * Var[Y j] := by + _ = ∑ j ∈ range (u (N - 1)), (∑ i ∈ range N with j < u i, ((u i : ℝ) ^ 2)⁻¹) * Var[Y j] := by simp_rw [mul_sum, sum_mul, sum_sigma'] refine sum_nbij' (fun p ↦ ⟨p.2, p.1⟩) (fun p ↦ ⟨p.2, p.1⟩) ?_ ?_ ?_ ?_ ?_ · simp only [mem_sigma, mem_range, filter_congr_decidable, mem_filter, and_imp, diff --git a/Mathlib/RingTheory/GradedAlgebra/Radical.lean b/Mathlib/RingTheory/GradedAlgebra/Radical.lean index 976b7748f818d..72e6ad6f51e2b 100644 --- a/Mathlib/RingTheory/GradedAlgebra/Radical.lean +++ b/Mathlib/RingTheory/GradedAlgebra/Radical.lean @@ -66,10 +66,10 @@ theorem Ideal.IsHomogeneous.isPrime_of_homogeneous_mem_or_mem {I : Ideal A} (hI This is a contradiction, because both `proj (max₁ + max₂) (x * y) ∈ I` and the sum on the right hand side is in `I` however `proj max₁ x * proj max₂ y` is not in `I`. -/ - set set₁ := (decompose 𝒜 x).support.filter (fun i => proj 𝒜 i x ∉ I) with set₁_eq - set set₂ := (decompose 𝒜 y).support.filter (fun i => proj 𝒜 i y ∉ I) with set₂_eq + set set₁ := {i ∈ (decompose 𝒜 x).support | proj 𝒜 i x ∉ I} with set₁_eq + set set₂ := {i ∈ (decompose 𝒜 y).support | proj 𝒜 i y ∉ I} with set₂_eq have nonempty : - ∀ x : A, x ∉ I → ((decompose 𝒜 x).support.filter (fun i => proj 𝒜 i x ∉ I)).Nonempty := by + ∀ x : A, x ∉ I → {i ∈ (decompose 𝒜 x).support | proj 𝒜 i x ∉ I}.Nonempty := by intro x hx rw [filter_nonempty_iff] contrapose! hx @@ -83,8 +83,8 @@ theorem Ideal.IsHomogeneous.isPrime_of_homogeneous_mem_or_mem {I : Ideal A} (hI replace hxy : proj 𝒜 (max₁ + max₂) (x * y) ∈ I := hI _ hxy have mem_I : proj 𝒜 max₁ x * proj 𝒜 max₂ y ∈ I := by set antidiag := - ((decompose 𝒜 x).support ×ˢ (decompose 𝒜 y).support).filter (fun z : ι × ι => - z.1 + z.2 = max₁ + max₂) with ha + {z ∈ (decompose 𝒜 x).support ×ˢ (decompose 𝒜 y).support | z.1 + z.2 = max₁ + max₂} + with ha have mem_antidiag : (max₁, max₂) ∈ antidiag := by simp only [antidiag, add_sum_erase, mem_filter, mem_product] exact ⟨⟨mem_of_mem_filter _ mem_max₁, mem_of_mem_filter _ mem_max₂⟩, trivial⟩ diff --git a/Mathlib/RingTheory/IntegralDomain.lean b/Mathlib/RingTheory/IntegralDomain.lean index 7126ac0db1145..b292e8b1c1ffc 100644 --- a/Mathlib/RingTheory/IntegralDomain.lean +++ b/Mathlib/RingTheory/IntegralDomain.lean @@ -109,14 +109,12 @@ end Ring variable [CommRing R] [IsDomain R] [Group G] --- Porting note: Finset doesn't seem to have `{g ∈ univ | g^n = g₀}` notation anymore, --- so we have to use `Finset.filter` instead theorem card_nthRoots_subgroup_units [Fintype G] [DecidableEq G] (f : G →* R) (hf : Injective f) {n : ℕ} (hn : 0 < n) (g₀ : G) : - Finset.card (Finset.univ.filter (fun g ↦ g^n = g₀)) ≤ Multiset.card (nthRoots n (f g₀)) := by + #{g | g ^ n = g₀} ≤ Multiset.card (nthRoots n (f g₀)) := by haveI : DecidableEq R := Classical.decEq _ calc - _ ≤ (nthRoots n (f g₀)).toFinset.card := card_le_card_of_injOn f (by aesop) hf.injOn + _ ≤ #(nthRoots n (f g₀)).toFinset := card_le_card_of_injOn f (by aesop) hf.injOn _ ≤ _ := (nthRoots n (f g₀)).toFinset_card_le /-- A finite subgroup of the unit group of an integral domain is cyclic. -/ @@ -194,12 +192,11 @@ theorem sum_hom_units_eq_zero (f : G →* R) (hf : f ≠ 1) : ∑ g : G, f g = 0 eq_comm] at hn replace hx1 : (x.val : R) - 1 ≠ 0 := -- Porting note: was `(x : R)` fun h => hx1 (Subtype.eq (Units.ext (sub_eq_zero.1 h))) - let c := (univ.filter fun g => f.toHomUnits g = 1).card + let c := #{g | f.toHomUnits g = 1} calc ∑ g : G, f g = ∑ g : G, (f.toHomUnits g : R) := rfl - _ = ∑ u ∈ univ.image f.toHomUnits, - (univ.filter fun g => f.toHomUnits g = u).card • (u : R) := - (sum_comp ((↑) : Rˣ → R) f.toHomUnits) + _ = ∑ u ∈ univ.image f.toHomUnits, #{g | f.toHomUnits g = u} • (u : R) := + sum_comp ((↑) : Rˣ → R) f.toHomUnits _ = ∑ u ∈ univ.image f.toHomUnits, c • (u : R) := (sum_congr rfl fun u hu => congr_arg₂ _ ?_ rfl) -- remaining goal 1, proven below @@ -211,7 +208,7 @@ theorem sum_hom_units_eq_zero (f : G →* R) (hf : f ≠ 1) : ∑ g : G, f g = 0 -- remaining goal 2, proven below _ = (0 : R) := smul_zero _ · -- remaining goal 1 - show (univ.filter fun g : G => f.toHomUnits g = u).card = c + show #{g : G | f.toHomUnits g = u} = c apply MonoidHom.card_fiber_eq_of_mem_range f.toHomUnits · simpa only [mem_image, mem_univ, true_and, Set.mem_range] using hu · exact ⟨1, f.toHomUnits.map_one⟩ diff --git a/Mathlib/RingTheory/MvPolynomial/Homogeneous.lean b/Mathlib/RingTheory/MvPolynomial/Homogeneous.lean index a5847ddf565cd..64a1bbcd81c41 100644 --- a/Mathlib/RingTheory/MvPolynomial/Homogeneous.lean +++ b/Mathlib/RingTheory/MvPolynomial/Homogeneous.lean @@ -457,8 +457,7 @@ theorem coeff_homogeneousComponent (d : σ →₀ ℕ) : convert coeff_weightedHomogeneousComponent n φ d theorem homogeneousComponent_apply : - homogeneousComponent n φ = - ∑ d ∈ φ.support.filter fun d => d.degree = n, monomial d (coeff d φ) := by + homogeneousComponent n φ = ∑ d ∈ φ.support with d.degree = n, monomial d (coeff d φ) := by simp_rw [degree_eq_weight_one] convert weightedHomogeneousComponent_apply n φ diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric/Defs.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric/Defs.lean index 2b4ebd09a9114..093790197e63a 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric/Defs.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric/Defs.lean @@ -198,7 +198,7 @@ theorem aeval_esymm_eq_multiset_esymm [Algebra R S] (n : ℕ) (f : σ → S) : /-- We can define `esymm σ R n` by summing over a subtype instead of over `powerset_len`. -/ theorem esymm_eq_sum_subtype (n : ℕ) : - esymm σ R n = ∑ t : { s : Finset σ // s.card = n }, ∏ i ∈ (t : Finset σ), X i := + esymm σ R n = ∑ t : {s : Finset σ // #s = n}, ∏ i ∈ (t : Finset σ), X i := sum_subtype _ (fun _ => mem_powersetCard_univ) _ /-- We can define `esymm σ R n` as a sum over explicit monomials -/ diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric/FundamentalTheorem.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric/FundamentalTheorem.lean index 8a1f28577910e..b822840d8c9ce 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric/FundamentalTheorem.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric/FundamentalTheorem.lean @@ -60,7 +60,7 @@ section accumulate /-- The `j`th entry of `accumulate n m t` is the sum of `t i` over all `i ≥ j`. -/ @[simps] def accumulate (n m : ℕ) : (Fin n → ℕ) →+ (Fin m → ℕ) where - toFun t j := ∑ i in univ.filter (fun i : Fin n ↦ (j : ℕ) ≤ i), t i + toFun t j := ∑ i : Fin n with j.val ≤ i.val, t i map_zero' := funext <| fun _ ↦ sum_eq_zero <| fun _ _ ↦ rfl map_add' t₁ t₂ := funext <| fun j ↦ by dsimp only; exact sum_add_distrib @@ -178,7 +178,7 @@ private lemma supDegree_monic_esymm [Nontrivial R] {i : ℕ} (him : i < m) : supDegree_single_ne_zero _ one_ne_zero, leadingCoeff_single toLex.injective] at this · exact mem_powersetCard.2 ⟨subset_univ _, Fin.card_Iic _⟩ intro t ht hne - have ht' : t.card = (Iic (⟨i, him⟩ : Fin m)).card := by + have ht' : #t = #(Iic (⟨i, him⟩ : Fin m)) := by rw [(mem_powersetCard.1 ht).2, Fin.card_Iic] simp_rw [← single_eq_monomial, supDegree_single_ne_zero _ one_ne_zero, ← Finsupp.indicator_eq_sum_single] diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric/NewtonIdentities.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric/NewtonIdentities.lean index 3aa09ead7b4ab..802d25319935b 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric/NewtonIdentities.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric/NewtonIdentities.lean @@ -87,15 +87,15 @@ private theorem pairMap_involutive : (pairMap σ).Involutive := by variable [Fintype σ] private def pairs (k : ℕ) : Finset (Finset σ × σ) := - univ.filter (fun t ↦ card t.fst ≤ k ∧ (card t.fst = k → t.snd ∈ t.fst)) + {t | #t.1 ≤ k ∧ (#t.1 = k → t.snd ∈ t.fst)} @[simp] private lemma mem_pairs (k : ℕ) (t : Finset σ × σ) : - t ∈ pairs σ k ↔ card t.fst ≤ k ∧ (card t.fst = k → t.snd ∈ t.fst) := by + t ∈ pairs σ k ↔ #t.1 ≤ k ∧ (#t.1 = k → t.snd ∈ t.fst) := by simp [pairs] private def weight (k : ℕ) (t : Finset σ × σ) : MvPolynomial σ R := - (-1) ^ card t.fst * ((∏ a ∈ t.fst, X a) * X t.snd ^ (k - card t.fst)) + (-1) ^ #t.1 * ((∏ a ∈ t.fst, X a) * X t.snd ^ (k - #t.1)) private theorem pairMap_mem_pairs {k : ℕ} (t : Finset σ × σ) (h : t ∈ pairs σ k) : pairMap σ t ∈ pairs σ k := by @@ -126,16 +126,16 @@ private theorem weight_add_weight_pairMap {k : ℕ} (t : Finset σ × σ) (h : t mul_assoc (∏ a ∈ erase t.fst t.snd, X a), card_erase_of_mem h1] nth_rewrite 1 [← pow_one (X t.snd)] simp only [← pow_add, add_comm] - have h3 : 1 ≤ card t.fst := lt_iff_add_one_le.mp (card_pos.mpr ⟨t.snd, h1⟩) - rw [← tsub_tsub_assoc h.left h3, ← neg_neg ((-1 : MvPolynomial σ R) ^ (card t.fst - 1)), - h2 (card t.fst - 1), Nat.sub_add_cancel h3] + have h3 : 1 ≤ #t.1 := lt_iff_add_one_le.mp (card_pos.mpr ⟨t.snd, h1⟩) + rw [← tsub_tsub_assoc h.left h3, ← neg_neg ((-1 : MvPolynomial σ R) ^ (#t.1 - 1)), + h2 (#t.1 - 1), Nat.sub_add_cancel h3] simp · rw [pairMap_of_snd_nmem_fst σ h1] simp only [mul_comm, mul_assoc (∏ a ∈ t.fst, X a), card_cons, prod_cons] nth_rewrite 2 [← pow_one (X t.snd)] simp only [← pow_add, ← Nat.add_sub_assoc (Nat.lt_of_le_of_ne h.left (mt h.right h1)), add_comm, Nat.succ_eq_add_one, Nat.add_sub_add_right] - rw [← neg_neg ((-1 : MvPolynomial σ R) ^ card t.fst), h2] + rw [← neg_neg ((-1 : MvPolynomial σ R) ^ #t.1), h2] simp private theorem weight_sum (k : ℕ) : ∑ t ∈ pairs σ k, weight σ R k t = 0 := @@ -145,45 +145,40 @@ private theorem weight_sum (k : ℕ) : ∑ t ∈ pairs σ k, weight σ R k t = 0 private theorem sum_filter_pairs_eq_sum_powersetCard_sum (k : ℕ) (f : Finset σ × σ → MvPolynomial σ R) : - (∑ t ∈ filter (fun t ↦ card t.fst = k) (pairs σ k), f t) = - ∑ A ∈ powersetCard k univ, (∑ j ∈ A, f (A, j)) := by + ∑ t ∈ pairs σ k with #t.1 = k, f t = ∑ A ∈ powersetCard k univ, ∑ j ∈ A, f (A, j) := by apply sum_finset_product aesop private theorem sum_filter_pairs_eq_sum_powersetCard_mem_filter_antidiagonal_sum (k : ℕ) (a : ℕ × ℕ) - (ha : a ∈ (antidiagonal k).filter (fun a ↦ a.fst < k)) (f : Finset σ × σ → MvPolynomial σ R) : - (∑ t ∈ filter (fun t ↦ card t.fst = a.fst) (pairs σ k), f t) = - ∑ A ∈ powersetCard a.fst univ, (∑ j, f (A, j)) := by + (ha : a ∈ {a ∈ antidiagonal k | a.fst < k}) (f : Finset σ × σ → MvPolynomial σ R) : + ∑ t ∈ pairs σ k with #t.1 = a.1, f t = ∑ A ∈ powersetCard a.1 univ, ∑ j, f (A, j) := by apply sum_finset_product simp only [mem_filter, mem_powersetCard_univ, mem_univ, and_true, and_iff_right_iff_imp] rintro p hp - have : card p.fst ≤ k := by apply le_of_lt; aesop + have : #p.fst ≤ k := by apply le_of_lt; aesop aesop private lemma filter_pairs_lt (k : ℕ) : - (pairs σ k).filter (fun (s, _) ↦ s.card < k) = + (pairs σ k).filter (fun (s, _) ↦ #s < k) = (range k).disjiUnion (powersetCard · univ) ((pairwise_disjoint_powersetCard _).set_pairwise _) ×ˢ univ := by ext; aesop (add unsafe le_of_lt) private theorem sum_filter_pairs_eq_sum_filter_antidiagonal_powersetCard_sum (k : ℕ) (f : Finset σ × σ → MvPolynomial σ R) : - ∑ t ∈ (pairs σ k).filter fun t ↦ card t.fst < k, f t = - ∑ a ∈ (antidiagonal k).filter fun a ↦ a.fst < k, - ∑ A ∈ powersetCard a.fst univ, ∑ j, f (A, j) := by + ∑ t ∈ pairs σ k with #t.1 < k, f t = + ∑ a ∈ antidiagonal k with a.fst < k, ∑ A ∈ powersetCard a.fst univ, ∑ j, f (A, j) := by rw [filter_pairs_lt, sum_product, sum_disjiUnion] refine sum_nbij' (fun n ↦ (n, k - n)) Prod.fst ?_ ?_ ?_ ?_ ?_ <;> simp (config := { contextual := true }) [@eq_comm _ _ k, Nat.add_sub_cancel', le_of_lt] private theorem disjoint_filter_pairs_lt_filter_pairs_eq (k : ℕ) : - Disjoint (filter (fun t ↦ card t.fst < k) (pairs σ k)) - (filter (fun t ↦ card t.fst = k) (pairs σ k)) := by + Disjoint {t ∈ pairs σ k | #t.1 < k} {t ∈ pairs σ k | #t.1 = k} := by rw [disjoint_filter] exact fun _ _ h1 h2 ↦ lt_irrefl _ (h2.symm.subst h1) private theorem disjUnion_filter_pairs_eq_pairs (k : ℕ) : - disjUnion (filter (fun t ↦ card t.fst < k) (pairs σ k)) - (filter (fun t ↦ card t.fst = k) (pairs σ k)) (disjoint_filter_pairs_lt_filter_pairs_eq σ k) = - pairs σ k := by + disjUnion {t ∈ pairs σ k | #t.1 < k} {t ∈ pairs σ k | #t.1 = k} + (disjoint_filter_pairs_lt_filter_pairs_eq σ k) = pairs σ k := by simp only [disjUnion_eq_union, Finset.ext_iff, pairs, filter_filter, mem_filter] intro a rw [← filter_or, mem_filter] @@ -200,7 +195,7 @@ private theorem esymm_summand_to_weight (k : ℕ) (A : Finset σ) (h : A ∈ pow simp [weight, mem_powersetCard_univ.mp h, mul_assoc] private theorem esymm_to_weight [DecidableEq σ] (k : ℕ) : k * esymm σ R k = - (-1) ^ k * ∑ t ∈ filter (fun t ↦ card t.fst = k) (pairs σ k), weight σ R k t := by + (-1) ^ k * ∑ t ∈ pairs σ k with #t.1 = k, weight σ R k t := by rw [esymm, sum_filter_pairs_eq_sum_powersetCard_sum σ R k (fun t ↦ weight σ R k t), sum_congr rfl (esymm_summand_to_weight σ R k), mul_comm (k : MvPolynomial σ R) ((-1) ^ k), ← mul_sum, ← mul_assoc, ← mul_assoc, ← pow_add, Even.neg_one_pow ⟨k, rfl⟩, one_mul] @@ -216,9 +211,8 @@ private theorem esymm_mul_psum_summand_to_weight (k : ℕ) (a : ℕ × ℕ) (ha rw [mem_powersetCard_univ.mp hs, ← mem_antidiagonal.mp ha, add_sub_self_left] private theorem esymm_mul_psum_to_weight [DecidableEq σ] (k : ℕ) : - ∑ a ∈ (antidiagonal k).filter (fun a ↦ a.fst < k), - (-1) ^ a.fst * esymm σ R a.fst * psum σ R a.snd = - ∑ t ∈ filter (fun t ↦ card t.fst < k) (pairs σ k), weight σ R k t := by + ∑ a ∈ antidiagonal k with a.fst < k, (-1) ^ a.fst * esymm σ R a.fst * psum σ R a.snd = + ∑ t ∈ pairs σ k with #t.1 < k, weight σ R k t := by rw [← sum_congr rfl (fun a ha ↦ esymm_mul_psum_summand_to_weight σ R k a (mem_filter.mp ha).left), sum_filter_pairs_eq_sum_filter_antidiagonal_powersetCard_sum σ R k] @@ -228,9 +222,9 @@ variable (σ : Type*) [Fintype σ] (R : Type*) [CommRing R] /-- **Newton's identities** give a recurrence relation for the kth elementary symmetric polynomial in terms of lower degree elementary symmetric polynomials and power sums. -/ -theorem mul_esymm_eq_sum (k : ℕ) : k * esymm σ R k = - (-1) ^ (k + 1) * ∑ a ∈ (antidiagonal k).filter (fun a ↦ a.fst < k), - (-1) ^ a.fst * esymm σ R a.fst * psum σ R a.snd := by +theorem mul_esymm_eq_sum (k : ℕ) : + k * esymm σ R k = (-1) ^ (k + 1) * + ∑ a ∈ antidiagonal k with a.1 < k, (-1) ^ a.1 * esymm σ R a.1 * psum σ R a.2 := by classical rw [NewtonIdentities.esymm_to_weight σ R k, NewtonIdentities.esymm_mul_psum_to_weight σ R k, eq_comm, ← sub_eq_zero, sub_eq_add_neg, neg_mul_eq_neg_mul, @@ -252,23 +246,22 @@ theorem sum_antidiagonal_card_esymm_psum_eq_zero : /-- 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 +theorem psum_eq_mul_esymm_sub_sum (k : ℕ) (h : 0 < k) : + psum σ R k = (-1) ^ (k + 1) * k * esymm σ R k - + ∑ a ∈ antidiagonal k with a.1 ∈ Set.Ioo 0 k, (-1) ^ a.fst * esymm σ R a.1 * psum σ R a.2 := by 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)) + rw [← (sum_filter_add_sum_filter_not {a ∈ antidiagonal k | a.fst < k} (fun a ↦ 0 < a.fst) (fun a ↦ (-1) ^ a.fst * esymm σ R a.fst * psum σ R a.snd))] at hesymm have sub_both_sides := congrArg (· - (-1 : MvPolynomial σ R) ^ (k + 1) * - ∑ a ∈ ((antidiagonal k).filter (fun a ↦ a.fst < k)).filter (fun a ↦ 0 < a.fst), + ∑ a ∈ {a ∈ antidiagonal k | a.fst < k} with 0 < a.fst, (-1) ^ a.fst * esymm σ R a.fst * psum σ R a.snd) hesymm simp only [left_distrib, add_sub_cancel_left] at sub_both_sides have sub_both_sides := congrArg ((-1 : MvPolynomial σ R) ^ (k + 1) * ·) sub_both_sides simp only [mul_sub_left_distrib, ← mul_assoc, ← pow_add, Even.neg_one_pow ⟨k + 1, rfl⟩, one_mul, not_le, lt_one_iff, filter_filter (fun a : ℕ × ℕ ↦ a.fst < k) (fun a ↦ ¬0 < a.fst)] at sub_both_sides - have : filter (fun a ↦ a.fst < k ∧ ¬0 < a.fst) (antidiagonal k) = {(0, k)} := by + have : {a ∈ antidiagonal k | a.fst < k ∧ ¬0 < a.fst} = {(0, k)} := by ext a rw [mem_filter, mem_antidiagonal, mem_singleton] refine ⟨?_, by rintro rfl; omega⟩ diff --git a/Mathlib/RingTheory/MvPolynomial/WeightedHomogeneous.lean b/Mathlib/RingTheory/MvPolynomial/WeightedHomogeneous.lean index d580a2fd428fc..18bd357e9e117 100644 --- a/Mathlib/RingTheory/MvPolynomial/WeightedHomogeneous.lean +++ b/Mathlib/RingTheory/MvPolynomial/WeightedHomogeneous.lean @@ -309,7 +309,7 @@ theorem coeff_weightedHomogeneousComponent [DecidableEq M] (d : σ →₀ ℕ) : theorem weightedHomogeneousComponent_apply [DecidableEq M] : weightedHomogeneousComponent w n φ = - ∑ d ∈ φ.support.filter fun d => weight w d = n, monomial d (coeff d φ) := + ∑ d ∈ φ.support with weight w d = n, monomial d (coeff d φ) := letI := Classical.decEq M Finsupp.filter_eq_sum (fun d : σ →₀ ℕ => weight w d = n) φ |>.trans <| by convert rfl diff --git a/Mathlib/RingTheory/MvPowerSeries/Basic.lean b/Mathlib/RingTheory/MvPowerSeries/Basic.lean index 3e649d881d0c8..ab05350f3f87a 100644 --- a/Mathlib/RingTheory/MvPowerSeries/Basic.lean +++ b/Mathlib/RingTheory/MvPowerSeries/Basic.lean @@ -671,7 +671,7 @@ theorem coeff_eq_zero_of_constantCoeff_nilpotent apply sum_eq_zero intro k hk rw [mem_finsuppAntidiag] at hk - set s := (range n).filter fun i ↦ k i = 0 with hs_def + set s := {i ∈ range n | k i = 0} with hs_def have hs : s ⊆ range n := filter_subset _ _ have hs' (i : ℕ) (hi : i ∈ s) : coeff R (k i) f = constantCoeff σ R f := by simp only [hs_def, mem_filter] at hi @@ -682,10 +682,10 @@ theorem coeff_eq_zero_of_constantCoeff_nilpotent rw [← prod_sdiff (s₁ := s) (filter_subset _ _)] apply mul_eq_zero_of_right rw [prod_congr rfl hs', prod_const] - suffices m ≤ s.card by + suffices m ≤ #s by obtain ⟨m', hm'⟩ := Nat.exists_eq_add_of_le this rw [hm', pow_add, hf, MulZeroClass.zero_mul] - rw [← Nat.add_le_add_iff_right, add_comm s.card, + rw [← Nat.add_le_add_iff_right, add_comm #s, Finset.card_sdiff_add_card_eq_card (filter_subset _ _), card_range] apply le_trans _ hn simp only [add_comm m, Nat.add_le_add_iff_right, ← hk.1, diff --git a/Mathlib/RingTheory/Polynomial/Vieta.lean b/Mathlib/RingTheory/Polynomial/Vieta.lean index 6ad52b3eb26ac..6df36b563f3c0 100644 --- a/Mathlib/RingTheory/Polynomial/Vieta.lean +++ b/Mathlib/RingTheory/Polynomial/Vieta.lean @@ -22,13 +22,10 @@ we derive `Polynomial.coeff_eq_esymm_roots_of_card`, the relationship between th the roots of `p` for a polynomial `p` that splits (i.e. having as many roots as its degree). -/ - -open Polynomial +open Finset Polynomial namespace Multiset -open Polynomial - section Semiring variable {R : Type*} [CommSemiring R] @@ -70,8 +67,8 @@ theorem prod_X_add_C_coeff' {σ} (s : Multiset σ) (r : σ → R) {k : ℕ} (h : (s.map fun i => X + C (r i)).prod.coeff k = (s.map r).esymm (Multiset.card s - k) := by erw [← map_map (fun r => X + C r) r, prod_X_add_C_coeff] <;> rw [s.card_map r]; assumption -theorem _root_.Finset.prod_X_add_C_coeff {σ} (s : Finset σ) (r : σ → R) {k : ℕ} (h : k ≤ s.card) : - (∏ i ∈ s, (X + C (r i))).coeff k = ∑ t ∈ s.powersetCard (s.card - k), ∏ i ∈ t, r i := by +theorem _root_.Finset.prod_X_add_C_coeff {σ} (s : Finset σ) (r : σ → R) {k : ℕ} (h : k ≤ #s) : + (∏ i ∈ s, (X + C (r i))).coeff k = ∑ t ∈ s.powersetCard (#s - k), ∏ i ∈ t, r i := by rw [Finset.prod, prod_X_add_C_coeff' _ r h, Finset.esymm_map_val] rfl diff --git a/Mathlib/RingTheory/Prime.lean b/Mathlib/RingTheory/Prime.lean index eba40fbe15828..6fb8a1e0598db 100644 --- a/Mathlib/RingTheory/Prime.lean +++ b/Mathlib/RingTheory/Prime.lean @@ -50,7 +50,7 @@ theorem mul_eq_mul_prime_pow {x y a p : R} {n : ℕ} (hp : Prime p) (hx : x * y rcases mul_eq_mul_prime_prod (fun _ _ ↦ hp) (show x * y = a * (range n).prod fun _ ↦ p by simpa) with ⟨t, u, b, c, htus, htu, rfl, rfl, rfl⟩ - exact ⟨t.card, u.card, b, c, by rw [← card_union_of_disjoint htu, htus, card_range], by simp⟩ + exact ⟨#t, #u, b, c, by rw [← card_union_of_disjoint htu, htus, card_range], by simp⟩ end CancelCommMonoidWithZero diff --git a/Mathlib/RingTheory/RootsOfUnity/Basic.lean b/Mathlib/RingTheory/RootsOfUnity/Basic.lean index a0b4d08c4b85d..b90cc6d9efb53 100644 --- a/Mathlib/RingTheory/RootsOfUnity/Basic.lean +++ b/Mathlib/RingTheory/RootsOfUnity/Basic.lean @@ -856,14 +856,14 @@ theorem nthRoots_one_nodup {ζ : R} {n : ℕ} (h : IsPrimitiveRoot ζ n) : @[simp] theorem card_nthRootsFinset {ζ : R} {n : ℕ} (h : IsPrimitiveRoot ζ n) : - (nthRootsFinset n R).card = n := by + #(nthRootsFinset n R) = n := by rw [nthRootsFinset, ← Multiset.toFinset_eq (nthRoots_one_nodup h), card_mk, h.card_nthRoots_one] open scoped Nat /-- If an integral domain has a primitive `k`-th root of unity, then it has `φ k` of them. -/ theorem card_primitiveRoots {ζ : R} {k : ℕ} (h : IsPrimitiveRoot ζ k) : - (primitiveRoots k R).card = φ k := by + #(primitiveRoots k R) = φ k := by by_cases h0 : k = 0 · simp [h0] symm diff --git a/Mathlib/RingTheory/RootsOfUnity/Lemmas.lean b/Mathlib/RingTheory/RootsOfUnity/Lemmas.lean index 33d3afe29785a..75a39566c1fe5 100644 --- a/Mathlib/RingTheory/RootsOfUnity/Lemmas.lean +++ b/Mathlib/RingTheory/RootsOfUnity/Lemmas.lean @@ -71,7 +71,7 @@ lemma self_sub_one_pow_dvd_order {k n : ℕ} (hn : k < n) {μ : R} (hμ : IsPrim rw [← this, mul_assoc, mul_assoc] congr 1 conv => enter [2, 2, 2]; rw [← card_range k] - rw [← prod_range_mul_prod_Ico _ (Nat.le_add_left k m), mul_comm _ (_ ^ card _), ← mul_assoc, + rw [← prod_range_mul_prod_Ico _ (Nat.le_add_left k m), mul_comm _ (_ ^ #_), ← mul_assoc, prod_mul_pow_card] conv => enter [2, 1, 2, j]; rw [← (Zdef _).2] diff --git a/Mathlib/Tactic/Positivity/Finset.lean b/Mathlib/Tactic/Positivity/Finset.lean index bef478ebd38c6..fdfd20e10b554 100644 --- a/Mathlib/Tactic/Positivity/Finset.lean +++ b/Mathlib/Tactic/Positivity/Finset.lean @@ -20,7 +20,7 @@ namespace Mathlib.Meta.Positivity open Qq Lean Meta Finset -/-- Extension for `Finset.card`. `s.card` is positive if `s` is nonempty. +/-- Extension for `Finset.card`. `#s` is positive if `s` is nonempty. It calls `Mathlib.Meta.proveFinsetNonempty` to attempt proving that the finset is nonempty. -/ @[positivity Finset.card _] @@ -42,7 +42,7 @@ def evalFintypeCard : PositivityExt where eval {u α} _ _ e := do return .positive q(@Fintype.card_pos $β $instβ $instβno) | _ => throwError "not Fintype.card" -/-- Extension for `Finset.dens`. `s.card` is positive if `s` is nonempty. +/-- Extension for `Finset.dens`. `s.dens` is positive if `s` is nonempty. It calls `Mathlib.Meta.proveFinsetNonempty` to attempt proving that the finset is nonempty. -/ @[positivity Finset.dens _] @@ -91,8 +91,8 @@ def evalFinsetSum : PositivityExt where eval {u α} zα pα e := do variable {α : Type*} {s : Finset α} -example : 0 ≤ s.card := by positivity -example (hs : s.Nonempty) : 0 < s.card := by positivity +example : 0 ≤ #s := by positivity +example (hs : s.Nonempty) : 0 < #s := by positivity variable [Fintype α] @@ -100,13 +100,13 @@ example : 0 ≤ Fintype.card α := by positivity example : 0 ≤ dens s := by positivity example (hs : s.Nonempty) : 0 < dens s := by positivity example (hs : s.Nonempty) : dens s ≠ 0 := by positivity -example [Nonempty α] : 0 < (univ : Finset α).card := by positivity +example [Nonempty α] : 0 < #(univ : Finset α) := by positivity example [Nonempty α] : 0 < Fintype.card α := by positivity example [Nonempty α] : 0 < dens (univ : Finset α) := by positivity example [Nonempty α] : dens (univ : Finset α) ≠ 0 := by positivity example {G : Type*} {A : Finset G} : - let f := fun _ : G ↦ 1; (∀ s, f s ^ 2 = 1) → 0 ≤ A.card := by + let f := fun _ : G ↦ 1; (∀ s, f s ^ 2 = 1) → 0 ≤ #A := by intros positivity -- Should succeed despite failing to prove `A` is nonempty. diff --git a/Mathlib/Topology/Algebra/InfiniteSum/Constructions.lean b/Mathlib/Topology/Algebra/InfiniteSum/Constructions.lean index b47e95b56f4ad..47d09ed95c689 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/Constructions.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/Constructions.lean @@ -86,7 +86,7 @@ theorem HasProd.sigma {γ : β → Type*} {f : (Σ b : β, γ b) → α} {g : β use u.image Sigma.fst, trivial intro bs hbs simp only [Set.mem_preimage, Finset.le_iff_subset] at hu - have : Tendsto (fun t : Finset (Σb, γ b) ↦ ∏ p ∈ t.filter fun p ↦ p.1 ∈ bs, f p) atTop + have : Tendsto (fun t : Finset (Σb, γ b) ↦ ∏ p ∈ t with p.1 ∈ bs, f p) atTop (𝓝 <| ∏ b ∈ bs, g b) := by simp only [← sigma_preimage_mk, prod_sigma] refine tendsto_finset_prod _ fun b _ ↦ ?_ @@ -161,17 +161,17 @@ theorem HasProd.of_sigma {γ : β → Type*} {f : (Σ b : β, γ b) → α} {g : obtain ⟨t0, st0, ht0⟩ : ∃ t0, ∏ i ∈ t0, g i ∈ v ∧ s.image Sigma.fst ⊆ t0 := by have A : ∀ᶠ t0 in (atTop : Filter (Finset β)), ∏ i ∈ t0, g i ∈ v := hg (v_open.mem_nhds hv) exact (A.and (Ici_mem_atTop _)).exists - have L : Tendsto (fun t : Finset (Σb, γ b) ↦ ∏ p ∈ t.filter fun p ↦ p.1 ∈ t0, f p) atTop + have L : Tendsto (fun t : Finset (Σb, γ b) ↦ ∏ p ∈ t with p.1 ∈ t0, f p) atTop (𝓝 <| ∏ b ∈ t0, g b) := by simp only [← sigma_preimage_mk, prod_sigma] refine tendsto_finset_prod _ fun b _ ↦ ?_ change Tendsto (fun t ↦ (fun t ↦ ∏ s ∈ t, f ⟨b, s⟩) (preimage t (Sigma.mk b) _)) atTop (𝓝 (g b)) exact (hf b).comp (tendsto_finset_preimage_atTop_atTop (sigma_mk_injective)) - have : ∃ t, ∏ p ∈ t.filter (fun p ↦ p.1 ∈ t0), f p ∈ v ∧ s ⊆ t := + have : ∃ t, ∏ p ∈ t with p.1 ∈ t0, f p ∈ v ∧ s ⊆ t := ((Tendsto.eventually_mem L (v_open.mem_nhds st0)).and (Ici_mem_atTop _)).exists obtain ⟨t, tv, st⟩ := this - refine ⟨t.filter (fun p ↦ p.1 ∈ t0), fun x hx ↦ ?_, vu tv⟩ + refine ⟨{p ∈ t | p.1 ∈ t0}, fun x hx ↦ ?_, vu tv⟩ simpa only [mem_filter, st hx, true_and] using ht0 (mem_image_of_mem Sigma.fst hx) variable [CompleteSpace α] diff --git a/Mathlib/Topology/Algebra/InfiniteSum/Group.lean b/Mathlib/Topology/Algebra/InfiniteSum/Group.lean index 9532cd4e53ab9..c4029fed54fca 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/Group.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/Group.lean @@ -241,9 +241,9 @@ theorem Multipliable.multipliable_of_eq_one_or_self (hf : Multipliable f) exact multipliable_iff_vanishing.2 fun e he ↦ let ⟨s, hs⟩ := multipliable_iff_vanishing.1 hf e he ⟨s, fun t ht ↦ - have eq : ∏ b ∈ t.filter fun b ↦ g b = f b, f b = ∏ b ∈ t, g b := + have eq : ∏ b ∈ t with g b = f b, f b = ∏ b ∈ t, g b := calc - ∏ b ∈ t.filter fun b ↦ g b = f b, f b = ∏ b ∈ t.filter fun b ↦ g b = f b, g b := + ∏ b ∈ t with g b = f b, f b = ∏ b ∈ t with g b = f b, g b := Finset.prod_congr rfl fun b hb ↦ (Finset.mem_filter.1 hb).2.symm _ = ∏ b ∈ t, g b := by {refine Finset.prod_subset (Finset.filter_subset _ _) ?_ diff --git a/Mathlib/Topology/Algebra/InfiniteSum/Order.lean b/Mathlib/Topology/Algebra/InfiniteSum/Order.lean index 88fad0d581136..612e335b22582 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/Order.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/Order.lean @@ -269,10 +269,10 @@ alias ⟨Summable.of_abs, Summable.abs⟩ := summable_abs_iff theorem Finite.of_summable_const [LinearOrderedAddCommGroup α] [TopologicalSpace α] [Archimedean α] [OrderClosedTopology α] {b : α} (hb : 0 < b) (hf : Summable fun _ : ι ↦ b) : Finite ι := by - have H : ∀ s : Finset ι, s.card • b ≤ ∑' _ : ι, b := fun s ↦ by + have H : ∀ s : Finset ι, #s • b ≤ ∑' _ : ι, b := fun s ↦ by simpa using sum_le_hasSum s (fun a _ ↦ hb.le) hf.hasSum obtain ⟨n, hn⟩ := Archimedean.arch (∑' _ : ι, b) hb - have : ∀ s : Finset ι, s.card ≤ n := fun s ↦ by + have : ∀ s : Finset ι, #s ≤ n := fun s ↦ by simpa [nsmul_le_nsmul_iff_left hb] using (H s).trans hn have : Fintype ι := fintypeOfFinsetCardLe n this infer_instance diff --git a/Mathlib/Topology/Instances/ENNReal.lean b/Mathlib/Topology/Instances/ENNReal.lean index bee9dc9f1a492..c227c3915781d 100644 --- a/Mathlib/Topology/Instances/ENNReal.lean +++ b/Mathlib/Topology/Instances/ENNReal.lean @@ -19,7 +19,7 @@ import Mathlib.Topology.Metrizable.Uniformity noncomputable section open Set Filter Metric Function -open scoped Topology ENNReal NNReal +open scoped Finset Topology ENNReal NNReal variable {α : Type*} {β : Type*} {γ : Type*} @@ -828,11 +828,11 @@ theorem finite_const_le_of_tsum_ne_top {ι : Type*} {a : ι → ℝ≥0∞} (tsu /-- Markov's inequality for `Finset.card` and `tsum` in `ℝ≥0∞`. -/ theorem finset_card_const_le_le_of_tsum_le {ι : Type*} {a : ι → ℝ≥0∞} {c : ℝ≥0∞} (c_ne_top : c ≠ ∞) (tsum_le_c : ∑' i, a i ≤ c) {ε : ℝ≥0∞} (ε_ne_zero : ε ≠ 0) : - ∃ hf : { i : ι | ε ≤ a i }.Finite, ↑hf.toFinset.card ≤ c / ε := by + ∃ hf : { i : ι | ε ≤ a i }.Finite, #hf.toFinset ≤ c / ε := by have hf : { i : ι | ε ≤ a i }.Finite := finite_const_le_of_tsum_ne_top (ne_top_of_le_ne_top c_ne_top tsum_le_c) ε_ne_zero refine ⟨hf, (ENNReal.le_div_iff_mul_le (.inl ε_ne_zero) (.inr c_ne_top)).2 ?_⟩ - calc ↑hf.toFinset.card * ε = ∑ _i ∈ hf.toFinset, ε := by rw [Finset.sum_const, nsmul_eq_mul] + calc #hf.toFinset * ε = ∑ _i ∈ hf.toFinset, ε := by rw [Finset.sum_const, nsmul_eq_mul] _ ≤ ∑ i ∈ hf.toFinset, a i := Finset.sum_le_sum fun i => hf.mem_toFinset.1 _ ≤ ∑' i, a i := ENNReal.sum_le_tsum _ _ ≤ c := tsum_le_c diff --git a/test/antidiagonal.lean b/test/antidiagonal.lean index c3e22116f5fe7..ea8da80cfaae9 100644 --- a/test/antidiagonal.lean +++ b/test/antidiagonal.lean @@ -14,9 +14,9 @@ section -- `antidiagonalTuple` is faster than `finAntidiagonal` by a small constant factor /-- info: 23426 -/ -#guard_msgs in #eval (finAntidiagonal 4 50).card +#guard_msgs in #eval #(finAntidiagonal 4 50) /-- info: 23426 -/ -#guard_msgs in #eval (Finset.Nat.antidiagonalTuple 4 50).card +#guard_msgs in #eval #(Finset.Nat.antidiagonalTuple 4 50) end diff --git a/test/positivity.lean b/test/positivity.lean index 07295214c3bbe..86a775d0e246c 100644 --- a/test/positivity.lean +++ b/test/positivity.lean @@ -290,8 +290,8 @@ example (n : ℕ+) : 0 < (↑n : ℕ) := by positivity example (n : ℕ) : 0 < n ! := by positivity example (n k : ℕ) : 0 < (n+1).ascFactorial k := by positivity --- example {α : Type _} (s : Finset α) (hs : s.Nonempty) : 0 < s.card := by positivity --- example {α : Type _} [Fintype α] [Nonempty α] : 0 < Fintype.card α := by positivity +example {α : Type _} (s : Finset α) (hs : s.Nonempty) : 0 < #s := by positivity +example {α : Type _} [Fintype α] [Nonempty α] : 0 < Fintype.card α := by positivity example {r : ℝ} : 0 < Real.exp r := by positivity From 7d7a2a8bb25e15c01279112b9f6e2a39ab74f7de Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Wed, 23 Oct 2024 02:43:32 +0000 Subject: [PATCH 314/505] chore(Algebra/Algebra/Hom): explain why `AlgHomClass.commutes` is not `@[simp]` (#18040) I came across this while cleaning up `simp`-related porting notes. It is tempting to move the `@[simp]` attribute from `AlgHom.commutes` to the generic lemma `AlgHomClass.commutes`, but this turns out to cause some slowdown. So replace the (slightly inaccurate) porting note with an up to date comment. --- Mathlib/Algebra/Algebra/Hom.lean | 4 +++- 1 file changed, 3 insertions(+), 1 deletion(-) diff --git a/Mathlib/Algebra/Algebra/Hom.lean b/Mathlib/Algebra/Algebra/Hom.lean index 5c5eac38616b5..b9f3122d4c8bd 100644 --- a/Mathlib/Algebra/Algebra/Hom.lean +++ b/Mathlib/Algebra/Algebra/Hom.lean @@ -48,7 +48,9 @@ class AlgHomClass (F : Type*) (R A B : outParam Type*) -- Porting note: `dangerousInstance` linter has become smarter about `outParam`s -- attribute [nolint dangerousInstance] AlgHomClass.toRingHomClass --- Porting note (#10618): simp can prove this +-- For now, don't replace `AlgHom.commutes` and `AlgHomClass.commutes` with the more generic lemma. +-- The file `Mathlib.NumberTheory.NumberField.CanonicalEmbedding.FundamentalCone` slows down by +-- 15% if we would do so (see benchmark on PR #18040). -- attribute [simp] AlgHomClass.commutes namespace AlgHomClass From b9ed435607e168cf3b58498b2041123940fa2460 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Wed, 23 Oct 2024 02:43:33 +0000 Subject: [PATCH 315/505] chore(scripts/create-adaptation-pr.sh): add message to ask for reviews (#18058) --- scripts/create-adaptation-pr.sh | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/scripts/create-adaptation-pr.sh b/scripts/create-adaptation-pr.sh index dedae3ef45b50..e383694363dac 100755 --- a/scripts/create-adaptation-pr.sh +++ b/scripts/create-adaptation-pr.sh @@ -204,7 +204,7 @@ if git diff --name-only bump/$BUMPVERSION bump/nightly-$NIGHTLYDATE | grep -q .; echo "### [auto/user] post a link to the PR on Zulip" zulip_title="#$pr_number adaptations for nightly-$NIGHTLYDATE" - zulip_body="> $pr_title #$pr_number" + zulip_body="> $pr_title #$pr_number\n\nPlease review this PR. At the end of the month this diff will land in \`master\`." echo "Post the link to the PR in a new thread on the #nightly-testing channel on Zulip" echo "Here is a suggested message:" From 1deca17a1b609ab594f02d2200998c396a52d1e7 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 23 Oct 2024 02:43:34 +0000 Subject: [PATCH 316/505] chore: Rename `QuotientMap` to `IsQuotientMap` (#18062) `Function.Embedding` is a type while `Embedding` is a proposition, and there are many other kinds of embeddings than topological embeddings. Hence this PR is a step towards 1. renaming `Embedding` to `IsEmbedding` and similarly for neighboring declarations (which `QuotientMap` is) 2. namespacing it inside `Topology` --- Mathlib/Analysis/Complex/ReImTopology.lean | 16 +++-- Mathlib/Analysis/Normed/Operator/Banach.lean | 7 +- Mathlib/Condensed/TopComparison.lean | 6 +- Mathlib/GroupTheory/QuotientGroup/Basic.lean | 2 +- Mathlib/RingTheory/Jacobson.lean | 4 +- .../Algebra/Constructions/DomMulAct.lean | 12 +++- Mathlib/Topology/Algebra/Group/Basic.lean | 3 +- Mathlib/Topology/Algebra/Group/Quotient.lean | 11 +-- .../Topology/Algebra/ProperAction/Basic.lean | 2 +- Mathlib/Topology/Algebra/Ring/Ideal.lean | 9 ++- .../Algebra/SeparationQuotient/Basic.lean | 10 +-- Mathlib/Topology/Bases.lean | 31 +++++--- .../Category/CompHausLike/EffectiveEpi.lean | 10 +-- .../Category/TopCat/EffectiveEpi.lean | 27 +++---- Mathlib/Topology/Clopen.lean | 5 +- Mathlib/Topology/CompactOpen.lean | 14 ++-- Mathlib/Topology/Connected/Clopen.lean | 21 ++++-- .../Connected/TotallyDisconnected.lean | 2 +- Mathlib/Topology/Constructions.lean | 35 +++++++--- Mathlib/Topology/ContinuousMap/Basic.lean | 16 ++--- Mathlib/Topology/ContinuousOn.lean | 7 +- Mathlib/Topology/Covering.lean | 7 +- Mathlib/Topology/Defs/Induced.lean | 9 ++- Mathlib/Topology/DiscreteQuotient.lean | 11 +-- Mathlib/Topology/FiberBundle/Basic.lean | 7 +- .../IsHomeomorphicTrivialBundle.lean | 9 ++- Mathlib/Topology/Homeomorph.lean | 20 ++++-- Mathlib/Topology/Inseparable.lean | 16 +++-- Mathlib/Topology/Instances/AddCircle.lean | 2 +- Mathlib/Topology/Maps/Basic.lean | 70 +++++++++++-------- Mathlib/Topology/Maps/OpenQuotient.lean | 23 ++++-- Mathlib/Topology/Order/ProjIcc.lean | 9 ++- Mathlib/Topology/Separation.lean | 11 +-- Mathlib/Topology/Sequences.lean | 7 +- scripts/no_lints_prime_decls.txt | 2 +- 35 files changed, 291 insertions(+), 162 deletions(-) diff --git a/Mathlib/Analysis/Complex/ReImTopology.lean b/Mathlib/Analysis/Complex/ReImTopology.lean index b9eac29a030c2..b157ee0e81a04 100644 --- a/Mathlib/Analysis/Complex/ReImTopology.lean +++ b/Mathlib/Analysis/Complex/ReImTopology.lean @@ -18,7 +18,7 @@ Each statement about `Complex.re` listed below has a counterpart about `Complex. * `Complex.isHomeomorphicTrivialFiberBundle_re`: `Complex.re` turns `ℂ` into a trivial topological fiber bundle over `ℝ`; -* `Complex.isOpenMap_re`, `Complex.quotientMap_re`: in particular, `Complex.re` is an open map +* `Complex.isOpenMap_re`, `Complex.isQuotientMap_re`: in particular, `Complex.re` is an open map and is a quotient map; * `Complex.interior_preimage_re`, `Complex.closure_preimage_re`, `Complex.frontier_preimage_re`: formulas for `interior (Complex.re ⁻¹' s)` etc; @@ -52,11 +52,17 @@ theorem isOpenMap_re : IsOpenMap re := theorem isOpenMap_im : IsOpenMap im := isHomeomorphicTrivialFiberBundle_im.isOpenMap_proj -theorem quotientMap_re : QuotientMap re := - isHomeomorphicTrivialFiberBundle_re.quotientMap_proj +theorem isQuotientMap_re : IsQuotientMap re := + isHomeomorphicTrivialFiberBundle_re.isQuotientMap_proj -theorem quotientMap_im : QuotientMap im := - isHomeomorphicTrivialFiberBundle_im.quotientMap_proj +@[deprecated (since := "2024-10-22")] +alias quotientMap_re := isQuotientMap_re + +theorem isQuotientMap_im : IsQuotientMap im := + isHomeomorphicTrivialFiberBundle_im.isQuotientMap_proj + +@[deprecated (since := "2024-10-22")] +alias quotientMap_im := isQuotientMap_im theorem interior_preimage_re (s : Set ℝ) : interior (re ⁻¹' s) = re ⁻¹' interior s := (isOpenMap_re.preimage_interior_eq_interior_preimage continuous_re _).symm diff --git a/Mathlib/Analysis/Normed/Operator/Banach.lean b/Mathlib/Analysis/Normed/Operator/Banach.lean index 33ad7e49abced..b98dc85ceda60 100644 --- a/Mathlib/Analysis/Normed/Operator/Banach.lean +++ b/Mathlib/Analysis/Normed/Operator/Banach.lean @@ -249,8 +249,11 @@ protected theorem isOpenMap (surj : Surjective f) : IsOpenMap f := by exact Set.mem_image_of_mem _ (hε this) -protected theorem quotientMap (surj : Surjective f) : QuotientMap f := - (f.isOpenMap surj).to_quotientMap f.continuous surj +theorem isQuotientMap (surj : Surjective f) : IsQuotientMap f := + (f.isOpenMap surj).isQuotientMap f.continuous surj + +@[deprecated (since := "2024-10-22")] +alias quotientMap := isQuotientMap end diff --git a/Mathlib/Condensed/TopComparison.lean b/Mathlib/Condensed/TopComparison.lean index bcb10b87c8ea9..d0e30e97e7708 100644 --- a/Mathlib/Condensed/TopComparison.lean +++ b/Mathlib/Condensed/TopComparison.lean @@ -31,7 +31,7 @@ variable {C : Type u} [Category.{v} C] (G : C ⥤ TopCat.{w}) (X : Type w') [TopologicalSpace X] /-- -An auxiliary lemma to that allows us to use `QuotientMap.lift` in the proof of +An auxiliary lemma to that allows us to use `IsQuotientMap.lift` in the proof of `equalizerCondition_yonedaPresheaf`. -/ theorem factorsThrough_of_pullbackCondition {Z B : C} {π : Z ⟶ B} [HasPullback π π] @@ -61,7 +61,7 @@ condition which is required to be a sheaf for the regular topology. -/ theorem equalizerCondition_yonedaPresheaf [∀ (Z B : C) (π : Z ⟶ B) [EffectiveEpi π], PreservesLimit (cospan π π) G] - (hq : ∀ (Z B : C) (π : Z ⟶ B) [EffectiveEpi π], QuotientMap (G.map π)) : + (hq : ∀ (Z B : C) (π : Z ⟶ B) [EffectiveEpi π], IsQuotientMap (G.map π)) : EqualizerCondition (yonedaPresheaf G X) := by apply EqualizerCondition.mk intro Z B π _ _ @@ -113,7 +113,7 @@ def TopCat.toSheafCompHausLike : apply (config := { allowSynthFailures := true }) equalizerCondition_yonedaPresheaf (CompHausLike.compHausLikeToTop.{u} P) X intro Z B π he - apply QuotientMap.of_surjective_continuous (hs _ he) π.continuous + apply IsQuotientMap.of_surjective_continuous (hs _ he) π.continuous /-- `TopCat.toSheafCompHausLike` yields a functor from `TopCat.{max u w}` to diff --git a/Mathlib/GroupTheory/QuotientGroup/Basic.lean b/Mathlib/GroupTheory/QuotientGroup/Basic.lean index 5afa39a81c9a4..207475f814a34 100644 --- a/Mathlib/GroupTheory/QuotientGroup/Basic.lean +++ b/Mathlib/GroupTheory/QuotientGroup/Basic.lean @@ -432,7 +432,7 @@ the equalities, since the type of `(A'.addSubgroupOf A : AddSubgroup A)` depends def equivQuotientSubgroupOfOfEq {A' A B' B : Subgroup G} [hAN : (A'.subgroupOf A).Normal] [hBN : (B'.subgroupOf B).Normal] (h' : A' = B') (h : A = B) : A ⧸ A'.subgroupOf A ≃* B ⧸ B'.subgroupOf B := - MonoidHom.toMulEquiv (quotientMapSubgroupOfOfLe h'.le h.le) (quotientMapSubgroupOfOfLe h'.ge h.ge) + (quotientMapSubgroupOfOfLe h'.le h.le).toMulEquiv (quotientMapSubgroupOfOfLe h'.ge h.ge) (by ext ⟨x, hx⟩; rfl) (by ext ⟨x, hx⟩; rfl) diff --git a/Mathlib/RingTheory/Jacobson.lean b/Mathlib/RingTheory/Jacobson.lean index 93ed9e32a751f..aaa724fa39606 100644 --- a/Mathlib/RingTheory/Jacobson.lean +++ b/Mathlib/RingTheory/Jacobson.lean @@ -285,8 +285,8 @@ theorem isIntegral_isLocalization_polynomial_quotient [IsLocalization.Away (pX.map (Quotient.mk (P.comap (C : R →+* R[X])))).leadingCoeff Rₘ] [Algebra (R[X] ⧸ P) Sₘ] [IsLocalization ((Submonoid.powers (pX.map (Quotient.mk (P.comap (C : R →+* R[X])))).leadingCoeff).map (quotientMap P C le_rfl) : Submonoid (R[X] ⧸ P)) Sₘ] : - (IsLocalization.map Sₘ (quotientMap P C le_rfl) (Submonoid.powers (pX.map (Quotient.mk (P.comap - (C : R →+* R[X])))).leadingCoeff).le_comap_map : Rₘ →+* Sₘ).IsIntegral := by + (IsLocalization.map Sₘ (quotientMap P C le_rfl) (Submonoid.powers (pX.map (Quotient.mk + (P.comap (C : R →+* R[X])))).leadingCoeff).le_comap_map : Rₘ →+* Sₘ).IsIntegral := by let P' : Ideal R := P.comap C let M : Submonoid (R ⧸ P') := Submonoid.powers (pX.map (Quotient.mk (P.comap (C : R →+* R[X])))).leadingCoeff diff --git a/Mathlib/Topology/Algebra/Constructions/DomMulAct.lean b/Mathlib/Topology/Algebra/Constructions/DomMulAct.lean index 575ad9565a29c..f8c0f313ada70 100644 --- a/Mathlib/Topology/Algebra/Constructions/DomMulAct.lean +++ b/Mathlib/Topology/Algebra/Constructions/DomMulAct.lean @@ -52,7 +52,10 @@ theorem coe_mkHomeomorph_symm : ⇑(mkHomeomorph : M ≃ₜ Mᵈᵐᵃ).symm = m @[to_additive] theorem isOpenEmbedding_mk : IsOpenEmbedding (@mk M) := mkHomeomorph.isOpenEmbedding @[to_additive] theorem isClosedEmbedding_mk : IsClosedEmbedding (@mk M) := mkHomeomorph.isClosedEmbedding -@[to_additive] theorem quotientMap_mk : QuotientMap (@mk M) := mkHomeomorph.quotientMap +@[to_additive] theorem isQuotientMap_mk : IsQuotientMap (@mk M) := mkHomeomorph.isQuotientMap + +@[deprecated (since := "2024-10-22")] +alias quotientMap_mk := isQuotientMap_mk @[deprecated (since := "2024-10-20")] alias closedEmbedding_mk := isClosedEmbedding_mk @@ -77,7 +80,10 @@ theorem isClosedEmbedding_mk_symm : IsClosedEmbedding (@mk M).symm := alias closedEmbedding_mk_symm := isClosedEmbedding_mk_symm @[to_additive] -theorem quotientMap_mk_symm : QuotientMap (@mk M).symm := mkHomeomorph.symm.quotientMap +theorem isQuotientMap_mk_symm : IsQuotientMap (@mk M).symm := mkHomeomorph.symm.isQuotientMap + +@[deprecated (since := "2024-10-22")] +alias quotientMap_mk_symm := isQuotientMap_mk_symm @[to_additive] instance instT0Space [T0Space M] : T0Space Mᵈᵐᵃ := embedding_mk_symm.t0Space @[to_additive] instance instT1Space [T1Space M] : T1Space Mᵈᵐᵃ := embedding_mk_symm.t1Space @@ -106,7 +112,7 @@ instance instDiscreteTopology [DiscreteTopology M] : DiscreteTopology Mᵈᵐᵃ @[to_additive] instance instSeparableSpace [SeparableSpace M] : SeparableSpace Mᵈᵐᵃ := - quotientMap_mk.separableSpace + isQuotientMap_mk.separableSpace @[to_additive] instance instFirstCountableTopology [FirstCountableTopology M] : FirstCountableTopology Mᵈᵐᵃ := diff --git a/Mathlib/Topology/Algebra/Group/Basic.lean b/Mathlib/Topology/Algebra/Group/Basic.lean index bb19df20f7401..49cb37c34c00c 100644 --- a/Mathlib/Topology/Algebra/Group/Basic.lean +++ b/Mathlib/Topology/Algebra/Group/Basic.lean @@ -1047,7 +1047,8 @@ theorem MulAction.isClosedMap_quotient [CompactSpace α] : letI := orbitRel α β IsClosedMap (Quotient.mk' : β → Quotient (orbitRel α β)) := by intro t ht - rw [← quotientMap_quotient_mk'.isClosed_preimage, MulAction.quotient_preimage_image_eq_union_mul] + rw [← isQuotientMap_quotient_mk'.isClosed_preimage, + MulAction.quotient_preimage_image_eq_union_mul] convert ht.smul_left_of_isCompact (isCompact_univ (X := α)) rw [← biUnion_univ, ← iUnion_smul_left_image] rfl diff --git a/Mathlib/Topology/Algebra/Group/Quotient.lean b/Mathlib/Topology/Algebra/Group/Quotient.lean index 08cb4e12c8933..ced9990e38fb3 100644 --- a/Mathlib/Topology/Algebra/Group/Quotient.lean +++ b/Mathlib/Topology/Algebra/Group/Quotient.lean @@ -30,8 +30,11 @@ instance [CompactSpace G] (N : Subgroup G) : CompactSpace (G ⧸ N) := Quotient.compactSpace @[to_additive] -theorem quotientMap_mk (N : Subgroup G) : QuotientMap (mk : G → G ⧸ N) := - quotientMap_quot_mk +theorem isQuotientMap_mk (N : Subgroup G) : IsQuotientMap (mk : G → G ⧸ N) := + isQuotientMap_quot_mk + +@[deprecated (since := "2024-10-22")] +alias quotientMap_mk := isQuotientMap_mk @[to_additive] theorem continuous_mk {N : Subgroup G} : Continuous (mk : G → G ⧸ N) := @@ -121,13 +124,13 @@ theorem _root_.topologicalGroup_quotient [N.Normal] : TopologicalGroup (G ⧸ N) theorem isClosedMap_coe {H : Subgroup G} (hH : IsCompact (H : Set G)) : IsClosedMap ((↑) : G → G ⧸ H) := by intro t ht - rw [← (quotientMap_mk H).isClosed_preimage, preimage_image_mk_eq_mul] + rw [← (isQuotientMap_mk H).isClosed_preimage, preimage_image_mk_eq_mul] exact ht.mul_right_of_isCompact hH @[to_additive] instance instT3Space [N.Normal] [hN : IsClosed (N : Set G)] : T3Space (G ⧸ N) := by rw [← QuotientGroup.ker_mk' N] at hN - haveI := TopologicalGroup.t1Space (G ⧸ N) ((quotientMap_mk N).isClosed_preimage.mp hN) + haveI := TopologicalGroup.t1Space (G ⧸ N) ((isQuotientMap_mk N).isClosed_preimage.mp hN) infer_instance end QuotientGroup diff --git a/Mathlib/Topology/Algebra/ProperAction/Basic.lean b/Mathlib/Topology/Algebra/ProperAction/Basic.lean index 8b179ad23ab0f..52a706f6ad93c 100644 --- a/Mathlib/Topology/Algebra/ProperAction/Basic.lean +++ b/Mathlib/Topology/Algebra/ProperAction/Basic.lean @@ -113,7 +113,7 @@ theorem t2Space_quotient_mulAction_of_properSMul [ProperSMul G X] : let π : X → Quotient R := Quotient.mk' have : IsOpenQuotientMap (Prod.map π π) := MulAction.isOpenQuotientMap_quotientMk.prodMap MulAction.isOpenQuotientMap_quotientMk - rw [← this.quotientMap.isClosed_preimage] + rw [← this.isQuotientMap.isClosed_preimage] convert ProperSMul.isProperMap_smul_pair.isClosedMap.isClosed_range · ext ⟨x₁, x₂⟩ simp only [mem_preimage, map_apply, mem_diagonal_iff, mem_range, Prod.mk.injEq, Prod.exists, diff --git a/Mathlib/Topology/Algebra/Ring/Ideal.lean b/Mathlib/Topology/Algebra/Ring/Ideal.lean index c72e3ee7afefa..28afb07beedb3 100644 --- a/Mathlib/Topology/Algebra/Ring/Ideal.lean +++ b/Mathlib/Topology/Algebra/Ring/Ideal.lean @@ -58,12 +58,15 @@ theorem QuotientRing.isOpenMap_coe : IsOpenMap (mk N) := theorem QuotientRing.isOpenQuotientMap_mk : IsOpenQuotientMap (mk N) := QuotientAddGroup.isOpenQuotientMap_mk -theorem QuotientRing.quotientMap_coe_coe : QuotientMap fun p : R × R => (mk N p.1, mk N p.2) := - ((isOpenQuotientMap_mk N).prodMap (isOpenQuotientMap_mk N)).quotientMap +theorem QuotientRing.isQuotientMap_coe_coe : IsQuotientMap fun p : R × R => (mk N p.1, mk N p.2) := + ((isOpenQuotientMap_mk N).prodMap (isOpenQuotientMap_mk N)).isQuotientMap + +@[deprecated (since := "2024-10-22")] +alias QuotientRing.quotientMap_coe_coe := QuotientRing.isQuotientMap_coe_coe instance topologicalRing_quotient : TopologicalRing (R ⧸ N) where __ := QuotientAddGroup.instTopologicalAddGroup _ - continuous_mul := (QuotientRing.quotientMap_coe_coe N).continuous_iff.2 <| + continuous_mul := (QuotientRing.isQuotientMap_coe_coe N).continuous_iff.2 <| continuous_quot_mk.comp continuous_mul end CommRing diff --git a/Mathlib/Topology/Algebra/SeparationQuotient/Basic.lean b/Mathlib/Topology/Algebra/SeparationQuotient/Basic.lean index 26a143140affc..1f2427e29b56f 100644 --- a/Mathlib/Topology/Algebra/SeparationQuotient/Basic.lean +++ b/Mathlib/Topology/Algebra/SeparationQuotient/Basic.lean @@ -36,7 +36,7 @@ theorem mk_smul (c : M) (x : X) : mk (c • x) = c • mk x := rfl @[to_additive] instance instContinuousConstSMul : ContinuousConstSMul M (SeparationQuotient X) where - continuous_const_smul c := quotientMap_mk.continuous_iff.2 <| + continuous_const_smul c := isQuotientMap_mk.continuous_iff.2 <| continuous_mk.comp <| continuous_const_smul c @[to_additive] @@ -67,7 +67,7 @@ end SMul instance instContinuousSMul {M X : Type*} [SMul M X] [TopologicalSpace M] [TopologicalSpace X] [ContinuousSMul M X] : ContinuousSMul M (SeparationQuotient X) where continuous_smul := by - rw [(IsOpenQuotientMap.id.prodMap isOpenQuotientMap_mk).quotientMap.continuous_iff] + rw [(IsOpenQuotientMap.id.prodMap isOpenQuotientMap_mk).isQuotientMap.continuous_iff] exact continuous_mk.comp continuous_smul instance instSMulZeroClass {M X : Type*} [Zero X] [SMulZeroClass M X] [TopologicalSpace X] @@ -93,7 +93,7 @@ theorem mk_mul [Mul M] [ContinuousMul M] (a b : M) : mk (a * b) = mk a * mk b := @[to_additive] instance instContinuousMul [Mul M] [ContinuousMul M] : ContinuousMul (SeparationQuotient M) where - continuous_mul := quotientMap_prodMap_mk.continuous_iff.2 <| continuous_mk.comp continuous_mul + continuous_mul := isQuotientMap_prodMap_mk.continuous_iff.2 <| continuous_mk.comp continuous_mul @[to_additive] instance instCommMagma [CommMagma M] [ContinuousMul M] : CommMagma (SeparationQuotient M) := @@ -154,7 +154,7 @@ theorem mk_inv [Inv G] [ContinuousInv G] (x : G) : mk x⁻¹ = (mk x)⁻¹ := rf @[to_additive] instance instContinuousInv [Inv G] [ContinuousInv G] : ContinuousInv (SeparationQuotient G) where - continuous_inv := quotientMap_mk.continuous_iff.2 <| continuous_mk.comp continuous_inv + continuous_inv := isQuotientMap_mk.continuous_iff.2 <| continuous_mk.comp continuous_inv @[to_additive] instance instInvolutiveInv [InvolutiveInv G] [ContinuousInv G] : @@ -175,7 +175,7 @@ theorem mk_div [Div G] [ContinuousDiv G] (x y : G) : mk (x / y) = mk x / mk y := @[to_additive] instance instContinuousDiv [Div G] [ContinuousDiv G] : ContinuousDiv (SeparationQuotient G) where - continuous_div' := quotientMap_prodMap_mk.continuous_iff.2 <| continuous_mk.comp continuous_div' + continuous_div' := isQuotientMap_prodMap_mk.continuous_iff.2 <| continuous_mk.comp continuous_div' instance instZSMul [AddGroup G] [TopologicalAddGroup G] : SMul ℤ (SeparationQuotient G) := inferInstance diff --git a/Mathlib/Topology/Bases.lean b/Mathlib/Topology/Bases.lean index e953974c7c1dd..796775ecb345e 100644 --- a/Mathlib/Topology/Bases.lean +++ b/Mathlib/Topology/Bases.lean @@ -352,10 +352,13 @@ protected theorem _root_.DenseRange.separableSpace [SeparableSpace α] [Topologi let ⟨s, s_cnt, s_dense⟩ := exists_countable_dense α ⟨⟨f '' s, Countable.image s_cnt f, h.dense_image h' s_dense⟩⟩ -theorem _root_.QuotientMap.separableSpace [SeparableSpace α] [TopologicalSpace β] {f : α → β} - (hf : QuotientMap f) : SeparableSpace β := +theorem _root_.IsQuotientMap.separableSpace [SeparableSpace α] [TopologicalSpace β] {f : α → β} + (hf : IsQuotientMap f) : SeparableSpace β := hf.surjective.denseRange.separableSpace hf.continuous +@[deprecated (since := "2024-10-22")] +alias _root_.QuotientMap.separableSpace := _root_.IsQuotientMap.separableSpace + /-- The product of two separable spaces is a separable space. -/ instance [TopologicalSpace β] [SeparableSpace α] [SeparableSpace β] : SeparableSpace (α × β) := by rcases exists_countable_dense α with ⟨s, hsc, hsd⟩ @@ -382,10 +385,10 @@ instance {ι : Type*} {X : ι → Type*} [∀ i, TopologicalSpace (X i)] [∀ i, exact hyu ⟨i, _⟩ instance [SeparableSpace α] {r : α → α → Prop} : SeparableSpace (Quot r) := - quotientMap_quot_mk.separableSpace + isQuotientMap_quot_mk.separableSpace instance [SeparableSpace α] {s : Setoid α} : SeparableSpace (Quotient s) := - quotientMap_quot_mk.separableSpace + isQuotientMap_quot_mk.separableSpace /-- A topological space with discrete topology is separable iff it is countable. -/ theorem separableSpace_iff_countable [DiscreteTopology α] : SeparableSpace α ↔ Countable α := by @@ -917,8 +920,8 @@ section Quotient variable {X : Type*} [TopologicalSpace X] {Y : Type*} [TopologicalSpace Y] {π : X → Y} /-- The image of a topological basis under an open quotient map is a topological basis. -/ -theorem IsTopologicalBasis.quotientMap {V : Set (Set X)} (hV : IsTopologicalBasis V) - (h' : QuotientMap π) (h : IsOpenMap π) : IsTopologicalBasis (Set.image π '' V) := by +theorem IsTopologicalBasis.isQuotientMap {V : Set (Set X)} (hV : IsTopologicalBasis V) + (h' : IsQuotientMap π) (h : IsOpenMap π) : IsTopologicalBasis (Set.image π '' V) := by apply isTopologicalBasis_of_isOpen_of_nhds · rintro - ⟨U, U_in_V, rfl⟩ apply h U (hV.isOpen U_in_V) @@ -931,13 +934,19 @@ theorem IsTopologicalBasis.quotientMap {V : Set (Set X)} (hV : IsTopologicalBasi have πZ_in_U : π '' Z ⊆ U := (Set.image_subset _ Z_in_W).trans (image_preimage_subset π U) exact ⟨π '' Z, ⟨Z, Z_in_V, rfl⟩, ⟨x, x_in_Z, rfl⟩, πZ_in_U⟩ +@[deprecated (since := "2024-10-22")] +alias IsTopologicalBasis.quotientMap := IsTopologicalBasis.isQuotientMap + /-- A second countable space is mapped by an open quotient map to a second countable space. -/ -theorem _root_.QuotientMap.secondCountableTopology [SecondCountableTopology X] (h' : QuotientMap π) - (h : IsOpenMap π) : SecondCountableTopology Y where +theorem _root_.IsQuotientMap.secondCountableTopology [SecondCountableTopology X] + (h' : IsQuotientMap π) (h : IsOpenMap π) : SecondCountableTopology Y where is_open_generated_countable := by obtain ⟨V, V_countable, -, V_generates⟩ := exists_countable_basis X exact ⟨Set.image π '' V, V_countable.image (Set.image π), - (V_generates.quotientMap h' h).eq_generateFrom⟩ + (V_generates.isQuotientMap h' h).eq_generateFrom⟩ + +@[deprecated (since := "2024-10-22")] +alias _root_.QuotientMap.secondCountableTopology := _root_.IsQuotientMap.secondCountableTopology variable {S : Setoid X} @@ -945,12 +954,12 @@ variable {S : Setoid X} theorem IsTopologicalBasis.quotient {V : Set (Set X)} (hV : IsTopologicalBasis V) (h : IsOpenMap (Quotient.mk' : X → Quotient S)) : IsTopologicalBasis (Set.image (Quotient.mk' : X → Quotient S) '' V) := - hV.quotientMap quotientMap_quotient_mk' h + hV.isQuotientMap isQuotientMap_quotient_mk' h /-- An open quotient of a second countable space is second countable. -/ theorem Quotient.secondCountableTopology [SecondCountableTopology X] (h : IsOpenMap (Quotient.mk' : X → Quotient S)) : SecondCountableTopology (Quotient S) := - quotientMap_quotient_mk'.secondCountableTopology h + isQuotientMap_quotient_mk'.secondCountableTopology h end Quotient diff --git a/Mathlib/Topology/Category/CompHausLike/EffectiveEpi.lean b/Mathlib/Topology/Category/CompHausLike/EffectiveEpi.lean index 3305414bc773f..cea29cf78b297 100644 --- a/Mathlib/Topology/Category/CompHausLike/EffectiveEpi.lean +++ b/Mathlib/Topology/Category/CompHausLike/EffectiveEpi.lean @@ -33,20 +33,20 @@ If `π` is a surjective morphism in `CompHausLike P`, then it is an effective ep noncomputable def effectiveEpiStruct {B X : CompHausLike P} (π : X ⟶ B) (hπ : Function.Surjective π) : EffectiveEpiStruct π where - desc e h := (QuotientMap.of_surjective_continuous hπ π.continuous).lift e fun a b hab ↦ + desc e h := (IsQuotientMap.of_surjective_continuous hπ π.continuous).lift e fun a b hab ↦ DFunLike.congr_fun (h ⟨fun _ ↦ a, continuous_const⟩ ⟨fun _ ↦ b, continuous_const⟩ (by ext; exact hab)) a - fac e h := ((QuotientMap.of_surjective_continuous hπ π.continuous).lift_comp e + fac e h := ((IsQuotientMap.of_surjective_continuous hπ π.continuous).lift_comp e fun a b hab ↦ DFunLike.congr_fun (h ⟨fun _ ↦ a, continuous_const⟩ ⟨fun _ ↦ b, continuous_const⟩ (by ext; exact hab)) a) uniq e h g hm := by - suffices g = (QuotientMap.of_surjective_continuous hπ π.continuous).liftEquiv ⟨e, + suffices g = (IsQuotientMap.of_surjective_continuous hπ π.continuous).liftEquiv ⟨e, fun a b hab ↦ DFunLike.congr_fun (h ⟨fun _ ↦ a, continuous_const⟩ ⟨fun _ ↦ b, continuous_const⟩ (by ext; exact hab)) a⟩ by assumption - rw [← Equiv.symm_apply_eq (QuotientMap.of_surjective_continuous hπ π.continuous).liftEquiv] + rw [← Equiv.symm_apply_eq (IsQuotientMap.of_surjective_continuous hπ π.continuous).liftEquiv] ext - simp only [QuotientMap.liftEquiv_symm_apply_coe, ContinuousMap.comp_apply, ← hm] + simp only [IsQuotientMap.liftEquiv_symm_apply_coe, ContinuousMap.comp_apply, ← hm] rfl theorem preregular [HasExplicitPullbacks P] diff --git a/Mathlib/Topology/Category/TopCat/EffectiveEpi.lean b/Mathlib/Topology/Category/TopCat/EffectiveEpi.lean index 35693a3a1dedc..75c863468db2e 100644 --- a/Mathlib/Topology/Category/TopCat/EffectiveEpi.lean +++ b/Mathlib/Topology/Category/TopCat/EffectiveEpi.lean @@ -9,7 +9,7 @@ import Mathlib.Topology.Category.TopCat.Limits.Pullbacks # Effective epimorphisms in `TopCat` -This file proves the result `TopCat.effectiveEpi_iff_quotientMap`: +This file proves the result `TopCat.effectiveEpi_iff_isQuotientMap`: The effective epimorphisms in `TopCat` are precisely the quotient maps. -/ @@ -22,22 +22,22 @@ namespace TopCat /-- Implementation: If `π` is a morphism in `TopCat` which is a quotient map, then it is an effective -epimorphism. The theorem `TopCat.effectiveEpi_iff_quotientMap` should be used instead of +epimorphism. The theorem `TopCat.effectiveEpi_iff_isQuotientMap` should be used instead of this definition. -/ noncomputable -def effectiveEpiStructOfQuotientMap {B X : TopCat.{u}} (π : X ⟶ B) (hπ : QuotientMap π) : +def effectiveEpiStructOfQuotientMap {B X : TopCat.{u}} (π : X ⟶ B) (hπ : IsQuotientMap π) : EffectiveEpiStruct π where - /- `QuotientMap.lift` gives the required morphism -/ + /- `IsQuotientMap.lift` gives the required morphism -/ desc e h := hπ.lift e fun a b hab ↦ DFunLike.congr_fun (h ⟨fun _ ↦ a, continuous_const⟩ ⟨fun _ ↦ b, continuous_const⟩ (by ext; exact hab)) a - /- `QuotientMap.lift_comp` gives the factorisation -/ + /- `IsQuotientMap.lift_comp` gives the factorisation -/ fac e h := (hπ.lift_comp e fun a b hab ↦ DFunLike.congr_fun (h ⟨fun _ ↦ a, continuous_const⟩ ⟨fun _ ↦ b, continuous_const⟩ (by ext; exact hab)) a) - /- Uniqueness follows from the fact that `QuotientMap.lift` is an equivalence (given by - `QuotientMap.liftEquiv`). -/ + /- Uniqueness follows from the fact that `IsQuotientMap.lift` is an equivalence (given by + `IsQuotientMap.liftEquiv`). -/ uniq e h g hm := by suffices g = hπ.liftEquiv ⟨e, fun a b hab ↦ DFunLike.congr_fun @@ -45,12 +45,12 @@ def effectiveEpiStructOfQuotientMap {B X : TopCat.{u}} (π : X ⟶ B) (hπ : Quo a⟩ by assumption rw [← Equiv.symm_apply_eq hπ.liftEquiv] ext - simp only [QuotientMap.liftEquiv_symm_apply_coe, ContinuousMap.comp_apply, ← hm] + simp only [IsQuotientMap.liftEquiv_symm_apply_coe, ContinuousMap.comp_apply, ← hm] rfl /-- The effective epimorphisms in `TopCat` are precisely the quotient maps. -/ -theorem effectiveEpi_iff_quotientMap {B X : TopCat.{u}} (π : X ⟶ B) : - EffectiveEpi π ↔ QuotientMap π := by +theorem effectiveEpi_iff_isQuotientMap {B X : TopCat.{u}} (π : X ⟶ B) : + EffectiveEpi π ↔ IsQuotientMap π := by /- The backward direction is given by `effectiveEpiStructOfQuotientMap` above. -/ refine ⟨fun _ ↦ ?_, fun hπ ↦ ⟨⟨effectiveEpiStructOfQuotientMap π hπ⟩⟩⟩ /- Since `TopCat` has pullbacks, `π` is in fact a `RegularEpi`. This means that it exhibits `B` as @@ -59,8 +59,8 @@ theorem effectiveEpi_iff_quotientMap {B X : TopCat.{u}} (π : X ⟶ B) : have hπ : RegularEpi π := inferInstance let F := parallelPair hπ.left hπ.right let i : B ≅ colimit F := hπ.isColimit.coconePointUniqueUpToIso (colimit.isColimit _) - suffices QuotientMap (homeoOfIso i ∘ π) by - simpa [← Function.comp_assoc] using (homeoOfIso i).symm.quotientMap.comp this + suffices IsQuotientMap (homeoOfIso i ∘ π) by + simpa [← Function.comp_assoc] using (homeoOfIso i).symm.isQuotientMap.comp this constructor /- Effective epimorphisms are epimorphisms and epimorphisms in `TopCat` are surjective. -/ · change Function.Surjective (π ≫ i.hom) @@ -73,4 +73,7 @@ theorem effectiveEpi_iff_quotientMap {B X : TopCat.{u}} (π : X ⟶ B) : rw [isOpen_coinduced (f := (homeoOfIso i ∘ π)), coequalizer_isOpen_iff _ U, ← this] rfl +@[deprecated (since := "2024-10-22")] +alias effectiveEpi_iff_quotientMap := effectiveEpi_iff_isQuotientMap + end TopCat diff --git a/Mathlib/Topology/Clopen.lean b/Mathlib/Topology/Clopen.lean index 8d4aa07d73434..a8d606274a879 100644 --- a/Mathlib/Topology/Clopen.lean +++ b/Mathlib/Topology/Clopen.lean @@ -120,10 +120,13 @@ theorem isClopen_range_sigmaMk {X : ι → Type*} [∀ i, TopologicalSpace (X i) IsClopen (Set.range (@Sigma.mk ι X i)) := ⟨isClosedEmbedding_sigmaMk.isClosed_range, isOpenEmbedding_sigmaMk.isOpen_range⟩ -protected theorem QuotientMap.isClopen_preimage {f : X → Y} (hf : QuotientMap f) {s : Set Y} : +protected theorem IsQuotientMap.isClopen_preimage {f : X → Y} (hf : IsQuotientMap f) {s : Set Y} : IsClopen (f ⁻¹' s) ↔ IsClopen s := and_congr hf.isClosed_preimage hf.isOpen_preimage +@[deprecated (since := "2024-10-22")] +alias QuotientMap.isClopen_preimage := IsQuotientMap.isClopen_preimage + theorem continuous_boolIndicator_iff_isClopen (U : Set X) : Continuous U.boolIndicator ↔ IsClopen U := by rw [continuous_bool_rng true, preimage_boolIndicator_true] diff --git a/Mathlib/Topology/CompactOpen.lean b/Mathlib/Topology/CompactOpen.lean index 9c199df740de2..4ff883b2e49b1 100644 --- a/Mathlib/Topology/CompactOpen.lean +++ b/Mathlib/Topology/CompactOpen.lean @@ -441,12 +441,12 @@ theorem continuousMapOfUnique_symm_apply [Unique X] (f : C(X, Y)) : end Homeomorph -section QuotientMap +section IsQuotientMap variable {X₀ X Y Z : Type*} [TopologicalSpace X₀] [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] [LocallyCompactSpace Y] {f : X₀ → X} -theorem QuotientMap.continuous_lift_prod_left (hf : QuotientMap f) {g : X × Y → Z} +theorem IsQuotientMap.continuous_lift_prod_left (hf : IsQuotientMap f) {g : X × Y → Z} (hg : Continuous fun p : X₀ × Y => g (f p.1, p.2)) : Continuous g := by let Gf : C(X₀, C(Y, Z)) := ContinuousMap.curry ⟨_, hg⟩ have h : ∀ x : X, Continuous fun y => g (x, y) := by @@ -459,11 +459,17 @@ theorem QuotientMap.continuous_lift_prod_left (hf : QuotientMap f) {g : X × Y exact Gf.continuous exact ContinuousMap.continuous_uncurry_of_continuous ⟨G, this⟩ -theorem QuotientMap.continuous_lift_prod_right (hf : QuotientMap f) {g : Y × X → Z} +@[deprecated (since := "2024-10-22")] +alias QuotientMap.continuous_lift_prod_left := IsQuotientMap.continuous_lift_prod_left + +theorem IsQuotientMap.continuous_lift_prod_right (hf : IsQuotientMap f) {g : Y × X → Z} (hg : Continuous fun p : Y × X₀ => g (p.1, f p.2)) : Continuous g := by have : Continuous fun p : X₀ × Y => g ((Prod.swap p).1, f (Prod.swap p).2) := hg.comp continuous_swap have : Continuous fun p : X₀ × Y => (g ∘ Prod.swap) (f p.1, p.2) := this exact (hf.continuous_lift_prod_left this).comp continuous_swap -end QuotientMap +@[deprecated (since := "2024-10-22")] +alias QuotientMap.continuous_lift_prod_right := IsQuotientMap.continuous_lift_prod_right + +end IsQuotientMap diff --git a/Mathlib/Topology/Connected/Clopen.lean b/Mathlib/Topology/Connected/Clopen.lean index cfba16bc1231f..4a2541134dba0 100644 --- a/Mathlib/Topology/Connected/Clopen.lean +++ b/Mathlib/Topology/Connected/Clopen.lean @@ -453,18 +453,24 @@ theorem preimage_connectedComponent_connected [TopologicalSpace β] {f : α → from (this.trans T₂_v.1).trans inter_subset_right exact preimage_mono h -theorem QuotientMap.preimage_connectedComponent [TopologicalSpace β] {f : α → β} - (hf : QuotientMap f) (h_fibers : ∀ y : β, IsConnected (f ⁻¹' {y})) (a : α) : +theorem IsQuotientMap.preimage_connectedComponent [TopologicalSpace β] {f : α → β} + (hf : IsQuotientMap f) (h_fibers : ∀ y : β, IsConnected (f ⁻¹' {y})) (a : α) : f ⁻¹' connectedComponent (f a) = connectedComponent a := ((preimage_connectedComponent_connected h_fibers (fun _ => hf.isClosed_preimage.symm) _).subset_connectedComponent mem_connectedComponent).antisymm (hf.continuous.mapsTo_connectedComponent a) -theorem QuotientMap.image_connectedComponent [TopologicalSpace β] {f : α → β} (hf : QuotientMap f) +@[deprecated (since := "2024-10-22")] +alias QuotientMap.preimage_connectedComponent := IsQuotientMap.preimage_connectedComponent + +lemma IsQuotientMap.image_connectedComponent [TopologicalSpace β] {f : α → β} (hf : IsQuotientMap f) (h_fibers : ∀ y : β, IsConnected (f ⁻¹' {y})) (a : α) : f '' connectedComponent a = connectedComponent (f a) := by rw [← hf.preimage_connectedComponent h_fibers, image_preimage_eq _ hf.surjective] +@[deprecated (since := "2024-10-22")] +alias QuotientMap.image_connectedComponent := IsQuotientMap.image_connectedComponent + end Preconnected section connectedComponentSetoid @@ -505,12 +511,15 @@ instance : TopologicalSpace (ConnectedComponents α) := theorem surjective_coe : Surjective (mk : α → ConnectedComponents α) := surjective_quot_mk _ -theorem quotientMap_coe : QuotientMap (mk : α → ConnectedComponents α) := - quotientMap_quot_mk +theorem isQuotientMap_coe : IsQuotientMap (mk : α → ConnectedComponents α) := + isQuotientMap_quot_mk + +@[deprecated (since := "2024-10-22")] +alias quotientMap_coe := isQuotientMap_coe @[continuity] theorem continuous_coe : Continuous (mk : α → ConnectedComponents α) := - quotientMap_coe.continuous + isQuotientMap_coe.continuous @[simp] theorem range_coe : range (mk : α → ConnectedComponents α) = univ := diff --git a/Mathlib/Topology/Connected/TotallyDisconnected.lean b/Mathlib/Topology/Connected/TotallyDisconnected.lean index b7f1ff02b4fb4..78fdd9f07db9d 100644 --- a/Mathlib/Topology/Connected/TotallyDisconnected.lean +++ b/Mathlib/Topology/Connected/TotallyDisconnected.lean @@ -268,7 +268,7 @@ instance ConnectedComponents.totallyDisconnectedSpace : TotallyDisconnectedSpace (ConnectedComponents α) := by rw [totallyDisconnectedSpace_iff_connectedComponent_singleton] refine ConnectedComponents.surjective_coe.forall.2 fun x => ?_ - rw [← ConnectedComponents.quotientMap_coe.image_connectedComponent, ← + rw [← ConnectedComponents.isQuotientMap_coe.image_connectedComponent, ← connectedComponents_preimage_singleton, image_preimage_eq _ ConnectedComponents.surjective_coe] refine ConnectedComponents.surjective_coe.forall.2 fun y => ?_ rw [connectedComponents_preimage_singleton] diff --git a/Mathlib/Topology/Constructions.lean b/Mathlib/Topology/Constructions.lean index 4a48f65c105f5..c2dffe31a13df 100644 --- a/Mathlib/Topology/Constructions.lean +++ b/Mathlib/Topology/Constructions.lean @@ -739,11 +739,17 @@ theorem isOpen_prod_iff' {s : Set X} {t : Set Y} : simp only [st.1.ne_empty, st.2.ne_empty, not_false_iff, or_false] at H exact H.1.prod H.2 -theorem quotientMap_fst [Nonempty Y] : QuotientMap (Prod.fst : X × Y → X) := - isOpenMap_fst.to_quotientMap continuous_fst Prod.fst_surjective +theorem isQuotientMap_fst [Nonempty Y] : IsQuotientMap (Prod.fst : X × Y → X) := + isOpenMap_fst.isQuotientMap continuous_fst Prod.fst_surjective -theorem quotientMap_snd [Nonempty X] : QuotientMap (Prod.snd : X × Y → Y) := - isOpenMap_snd.to_quotientMap continuous_snd Prod.snd_surjective +@[deprecated (since := "2024-10-22")] +alias quotientMap_fst := isQuotientMap_fst + +theorem isQuotientMap_snd [Nonempty X] : IsQuotientMap (Prod.snd : X × Y → Y) := + isOpenMap_snd.isQuotientMap continuous_snd Prod.snd_surjective + +@[deprecated (since := "2024-10-22")] +alias quotientMap_snd := isQuotientMap_snd theorem closure_prod_eq {s : Set X} {t : Set Y} : closure (s ×ˢ t) = closure s ×ˢ closure t := ext fun ⟨a, b⟩ => by @@ -1143,13 +1149,16 @@ theorem DiscreteTopology.preimage_of_continuous_injective {X Y : Type*} [Topolog /-- If `f : X → Y` is a quotient map, then its restriction to the preimage of an open set is a quotient map too. -/ -theorem QuotientMap.restrictPreimage_isOpen {f : X → Y} (hf : QuotientMap f) - {s : Set Y} (hs : IsOpen s) : QuotientMap (s.restrictPreimage f) := by - refine quotientMap_iff.2 ⟨hf.surjective.restrictPreimage _, fun U ↦ ?_⟩ +theorem IsQuotientMap.restrictPreimage_isOpen {f : X → Y} (hf : IsQuotientMap f) + {s : Set Y} (hs : IsOpen s) : IsQuotientMap (s.restrictPreimage f) := by + refine isQuotientMap_iff.2 ⟨hf.surjective.restrictPreimage _, fun U ↦ ?_⟩ rw [hs.isOpenEmbedding_subtypeVal.open_iff_image_open, ← hf.isOpen_preimage, (hs.preimage hf.continuous).isOpenEmbedding_subtypeVal.open_iff_image_open, image_val_preimage_restrictPreimage] +@[deprecated (since := "2024-10-22")] +alias QuotientMap.restrictPreimage_isOpen := IsQuotientMap.restrictPreimage_isOpen + open scoped Set.Notation in lemma isClosed_preimage_val {s t : Set X} : IsClosed (s ↓∩ t) ↔ s ∩ closure (s ∩ t) ⊆ t := by rw [← closure_eq_iff_isClosed, embedding_subtype_val.closure_eq_preimage_closure_image, @@ -1171,9 +1180,12 @@ section Quotient variable [TopologicalSpace X] [TopologicalSpace Y] variable {r : X → X → Prop} {s : Setoid X} -theorem quotientMap_quot_mk : QuotientMap (@Quot.mk X r) := +theorem isQuotientMap_quot_mk : IsQuotientMap (@Quot.mk X r) := ⟨Quot.exists_rep, rfl⟩ +@[deprecated (since := "2024-10-22")] +alias quotientMap_quot_mk := isQuotientMap_quot_mk + @[continuity, fun_prop] theorem continuous_quot_mk : Continuous (@Quot.mk X r) := continuous_coinduced_rng @@ -1183,8 +1195,11 @@ theorem continuous_quot_lift {f : X → Y} (hr : ∀ a b, r a b → f a = f b) ( Continuous (Quot.lift f hr : Quot r → Y) := continuous_coinduced_dom.2 h -theorem quotientMap_quotient_mk' : QuotientMap (@Quotient.mk' X s) := - quotientMap_quot_mk +theorem isQuotientMap_quotient_mk' : IsQuotientMap (@Quotient.mk' X s) := + isQuotientMap_quot_mk + +@[deprecated (since := "2024-10-22")] +alias quotientMap_quotient_mk' := isQuotientMap_quotient_mk' theorem continuous_quotient_mk' : Continuous (@Quotient.mk' X s) := continuous_coinduced_rng diff --git a/Mathlib/Topology/ContinuousMap/Basic.lean b/Mathlib/Topology/ContinuousMap/Basic.lean index 69e4f1e0ee255..852cf38dfab8f 100644 --- a/Mathlib/Topology/ContinuousMap/Basic.lean +++ b/Mathlib/Topology/ContinuousMap/Basic.lean @@ -363,19 +363,19 @@ variable {X Y Z : Type*} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalS def Function.RightInverse.homeomorph {f' : C(Y, X)} (hf : Function.RightInverse f' f) : Quotient (Setoid.ker f) ≃ₜ Y where toEquiv := Setoid.quotientKerEquivOfRightInverse _ _ hf - continuous_toFun := quotientMap_quot_mk.continuous_iff.mpr (map_continuous f) + continuous_toFun := isQuotientMap_quot_mk.continuous_iff.mpr (map_continuous f) continuous_invFun := continuous_quotient_mk'.comp (map_continuous f') -namespace QuotientMap +namespace IsQuotientMap /-- The homeomorphism from the quotient of a quotient map to its codomain. This is `Setoid.quotientKerEquivOfSurjective` as a homeomorphism. -/ @[simps!] -noncomputable def homeomorph (hf : QuotientMap f) : Quotient (Setoid.ker f) ≃ₜ Y where +noncomputable def homeomorph (hf : IsQuotientMap f) : Quotient (Setoid.ker f) ≃ₜ Y where toEquiv := Setoid.quotientKerEquivOfSurjective _ hf.surjective - continuous_toFun := quotientMap_quot_mk.continuous_iff.mpr hf.continuous + continuous_toFun := isQuotientMap_quot_mk.continuous_iff.mpr hf.continuous continuous_invFun := by rw [hf.continuous_iff] convert continuous_quotient_mk' @@ -384,7 +384,7 @@ noncomputable def homeomorph (hf : QuotientMap f) : Quotient (Setoid.ker f) ≃ (Setoid.quotientKerEquivOfSurjective f hf.surjective).symm_apply_eq] rfl -variable (hf : QuotientMap f) (g : C(X, Z)) (h : Function.FactorsThrough g f) +variable (hf : IsQuotientMap f) (g : C(X, Z)) (h : Function.FactorsThrough g f) /-- Descend a continuous map, which is constant on the fibres, along a quotient map. -/ @[simps] @@ -394,7 +394,7 @@ noncomputable def lift : C(Y, Z) where continuous_toFun := Continuous.comp (continuous_quot_lift _ g.2) (Homeomorph.continuous _) /-- -The obvious triangle induced by `QuotientMap.lift` commutes: +The obvious triangle induced by `IsQuotientMap.lift` commutes: ``` g X --→ Z @@ -409,7 +409,7 @@ theorem lift_comp : (hf.lift g h).comp f = g := by ext simpa using h (Function.rightInverse_surjInv _ _) -/-- `QuotientMap.lift` as an equivalence. -/ +/-- `IsQuotientMap.lift` as an equivalence. -/ @[simps] noncomputable def liftEquiv : { g : C(X, Z) // Function.FactorsThrough g f} ≃ C(Y, Z) where toFun g := hf.lift g g.prop @@ -420,7 +420,7 @@ noncomputable def liftEquiv : { g : C(X, Z) // Function.FactorsThrough g f} ≃ ext a simpa using congrArg g (Function.rightInverse_surjInv hf.surjective a) -end QuotientMap +end IsQuotientMap end Lift diff --git a/Mathlib/Topology/ContinuousOn.lean b/Mathlib/Topology/ContinuousOn.lean index ef53ffeb6a397..9cd83acf60189 100644 --- a/Mathlib/Topology/ContinuousOn.lean +++ b/Mathlib/Topology/ContinuousOn.lean @@ -1090,11 +1090,14 @@ theorem IsOpenEmbedding.map_nhdsWithin_preimage_eq {f : α → β} (hf : IsOpenE @[deprecated (since := "2024-10-18")] alias OpenEmbedding.map_nhdsWithin_preimage_eq := IsOpenEmbedding.map_nhdsWithin_preimage_eq -theorem QuotientMap.continuousOn_isOpen_iff {f : α → β} {g : β → γ} (h : QuotientMap f) {s : Set β} - (hs : IsOpen s) : ContinuousOn g s ↔ ContinuousOn (g ∘ f) (f ⁻¹' s) := by +theorem IsQuotientMap.continuousOn_isOpen_iff {f : α → β} {g : β → γ} (h : IsQuotientMap f) + {s : Set β} (hs : IsOpen s) : ContinuousOn g s ↔ ContinuousOn (g ∘ f) (f ⁻¹' s) := by simp only [continuousOn_iff_continuous_restrict, (h.restrictPreimage_isOpen hs).continuous_iff] rfl +@[deprecated (since := "2024-10-22")] +alias QuotientMap.continuousOn_isOpen_iff := IsQuotientMap.continuousOn_isOpen_iff + theorem IsOpenMap.continuousOn_image_of_leftInvOn {f : α → β} {s : Set α} (h : IsOpenMap (s.restrict f)) {finv : β → α} (hleft : LeftInvOn finv f s) : ContinuousOn finv (f '' s) := by diff --git a/Mathlib/Topology/Covering.lean b/Mathlib/Topology/Covering.lean index 97aae6bbbf86b..10e634e668440 100644 --- a/Mathlib/Topology/Covering.lean +++ b/Mathlib/Topology/Covering.lean @@ -152,8 +152,11 @@ protected theorem isLocalHomeomorph : IsLocalHomeomorph f := protected theorem isOpenMap : IsOpenMap f := hf.isLocalHomeomorph.isOpenMap -protected theorem quotientMap (hf' : Function.Surjective f) : QuotientMap f := - hf.isOpenMap.to_quotientMap hf.continuous hf' +theorem isQuotientMap (hf' : Function.Surjective f) : IsQuotientMap f := + hf.isOpenMap.isQuotientMap hf.continuous hf' + +@[deprecated (since := "2024-10-22")] +alias quotientMap := isQuotientMap protected theorem isSeparatedMap : IsSeparatedMap f := fun e₁ e₂ he hne ↦ by diff --git a/Mathlib/Topology/Defs/Induced.lean b/Mathlib/Topology/Defs/Induced.lean index 0e1df6198a3bd..e76112cfb337f 100644 --- a/Mathlib/Topology/Defs/Induced.lean +++ b/Mathlib/Topology/Defs/Induced.lean @@ -36,7 +36,7 @@ as well as topology inducing maps, topological embeddings, and quotient maps. if it is an embedding and its range is open. An open embedding is an open map. -* `QuotientMap`: a map `f : X → Y` is a *quotient map*, +* `IsQuotientMap`: a map `f : X → Y` is a *quotient map*, if it is surjective and the topology on the codomain is equal to the coinduced topology. -/ @@ -130,8 +130,11 @@ alias ClosedEmbedding := IsClosedEmbedding /-- A function between topological spaces is a quotient map if it is surjective, and for all `s : Set Y`, `s` is open iff its preimage is an open set. -/ -@[mk_iff quotientMap_iff'] -structure QuotientMap {X : Type*} {Y : Type*} [tX : TopologicalSpace X] [tY : TopologicalSpace Y] +@[mk_iff isQuotientMap_iff'] +structure IsQuotientMap {X : Type*} {Y : Type*} [tX : TopologicalSpace X] [tY : TopologicalSpace Y] (f : X → Y) : Prop where surjective : Function.Surjective f eq_coinduced : tY = tX.coinduced f + +@[deprecated (since := "2024-10-22")] +alias QuotientMap := IsQuotientMap diff --git a/Mathlib/Topology/DiscreteQuotient.lean b/Mathlib/Topology/DiscreteQuotient.lean index bad892ec52aeb..4c20019504bc8 100644 --- a/Mathlib/Topology/DiscreteQuotient.lean +++ b/Mathlib/Topology/DiscreteQuotient.lean @@ -107,15 +107,18 @@ theorem fiber_eq (x : X) : S.proj ⁻¹' {S.proj x} = setOf (S.toSetoid x) := theorem proj_surjective : Function.Surjective S.proj := Quotient.surjective_Quotient_mk'' -theorem proj_quotientMap : QuotientMap S.proj := - quotientMap_quot_mk +theorem proj_isQuotientMap : IsQuotientMap S.proj := + isQuotientMap_quot_mk + +@[deprecated (since := "2024-10-22")] +alias proj_quotientMap := proj_isQuotientMap theorem proj_continuous : Continuous S.proj := - S.proj_quotientMap.continuous + S.proj_isQuotientMap.continuous instance : DiscreteTopology S := singletons_open_iff_discrete.1 <| S.proj_surjective.forall.2 fun x => by - rw [← S.proj_quotientMap.isOpen_preimage, fiber_eq] + rw [← S.proj_isQuotientMap.isOpen_preimage, fiber_eq] exact S.isOpen_setOf_rel _ theorem proj_isLocallyConstant : IsLocallyConstant S.proj := diff --git a/Mathlib/Topology/FiberBundle/Basic.lean b/Mathlib/Topology/FiberBundle/Basic.lean index 3fb66e7803f27..7fc7b01fd62b1 100644 --- a/Mathlib/Topology/FiberBundle/Basic.lean +++ b/Mathlib/Topology/FiberBundle/Basic.lean @@ -252,8 +252,11 @@ theorem surjective_proj [Nonempty F] : Function.Surjective (π F E) := fun b => /-- The projection from a fiber bundle with a nonempty fiber to its base is a quotient map. -/ -theorem quotientMap_proj [Nonempty F] : QuotientMap (π F E) := - (isOpenMap_proj F E).to_quotientMap (continuous_proj F E) (surjective_proj F E) +theorem isQuotientMap_proj [Nonempty F] : IsQuotientMap (π F E) := + (isOpenMap_proj F E).isQuotientMap (continuous_proj F E) (surjective_proj F E) + +@[deprecated (since := "2024-10-22")] +alias quotientMap_proj := isQuotientMap_proj theorem continuous_totalSpaceMk (x : B) : Continuous (@TotalSpace.mk B F E x) := (totalSpaceMk_inducing F E x).continuous diff --git a/Mathlib/Topology/FiberBundle/IsHomeomorphicTrivialBundle.lean b/Mathlib/Topology/FiberBundle/IsHomeomorphicTrivialBundle.lean index 865f13eb7f366..cc38260442348 100644 --- a/Mathlib/Topology/FiberBundle/IsHomeomorphicTrivialBundle.lean +++ b/Mathlib/Topology/FiberBundle/IsHomeomorphicTrivialBundle.lean @@ -52,9 +52,12 @@ protected theorem isOpenMap_proj (h : IsHomeomorphicTrivialFiberBundle F proj) : obtain ⟨e, rfl⟩ := h.proj_eq; exact isOpenMap_fst.comp e.isOpenMap /-- The projection from a trivial fiber bundle to its base is open. -/ -protected theorem quotientMap_proj [Nonempty F] (h : IsHomeomorphicTrivialFiberBundle F proj) : - QuotientMap proj := - h.isOpenMap_proj.to_quotientMap h.continuous_proj h.surjective_proj +protected theorem isQuotientMap_proj [Nonempty F] (h : IsHomeomorphicTrivialFiberBundle F proj) : + IsQuotientMap proj := + h.isOpenMap_proj.isQuotientMap h.continuous_proj h.surjective_proj + +@[deprecated (since := "2024-10-22")] +alias quotientMap_proj := IsHomeomorphicTrivialFiberBundle.isQuotientMap_proj end IsHomeomorphicTrivialFiberBundle diff --git a/Mathlib/Topology/Homeomorph.lean b/Mathlib/Topology/Homeomorph.lean index d1f3d9a2db354..28a19e26b5b3d 100644 --- a/Mathlib/Topology/Homeomorph.lean +++ b/Mathlib/Topology/Homeomorph.lean @@ -210,12 +210,15 @@ protected theorem inducing (h : X ≃ₜ Y) : Inducing h := theorem induced_eq (h : X ≃ₜ Y) : TopologicalSpace.induced h ‹_› = ‹_› := h.inducing.1.symm -protected theorem quotientMap (h : X ≃ₜ Y) : QuotientMap h := - QuotientMap.of_quotientMap_compose h.symm.continuous h.continuous <| by - simp only [self_comp_symm, QuotientMap.id] +theorem isQuotientMap (h : X ≃ₜ Y) : IsQuotientMap h := + IsQuotientMap.of_comp h.symm.continuous h.continuous <| by + simp only [self_comp_symm, IsQuotientMap.id] + +@[deprecated (since := "2024-10-22")] +alias quotientMap := isQuotientMap theorem coinduced_eq (h : X ≃ₜ Y) : TopologicalSpace.coinduced h ‹_› = ‹_› := - h.quotientMap.2.symm + h.isQuotientMap.2.symm protected theorem embedding (h : X ≃ₜ Y) : Embedding h := ⟨h.inducing, h.injective⟩ @@ -314,7 +317,7 @@ alias denseEmbedding := isDenseEmbedding @[simp] theorem isOpen_preimage (h : X ≃ₜ Y) {s : Set Y} : IsOpen (h ⁻¹' s) ↔ IsOpen s := - h.quotientMap.isOpen_preimage + h.isQuotientMap.isOpen_preimage @[simp] theorem isOpen_image (h : X ≃ₜ Y) {s : Set X} : IsOpen (h '' s) ↔ IsOpen s := by @@ -450,7 +453,7 @@ theorem comp_continuous_iff (h : X ≃ₜ Y) {f : Z → X} : Continuous (h ∘ f @[simp] theorem comp_continuous_iff' (h : X ≃ₜ Y) {f : Y → Z} : Continuous (f ∘ h) ↔ Continuous f := - h.quotientMap.continuous_iff.symm + h.isQuotientMap.continuous_iff.symm theorem comp_continuousAt_iff (h : X ≃ₜ Y) (f : Z → X) (z : Z) : ContinuousAt (h ∘ f) z ↔ ContinuousAt f z := @@ -912,12 +915,15 @@ noncomputable def homeomorph : X ≃ₜ Y where protected lemma isClosedMap : IsClosedMap f := (hf.homeomorph f).isClosedMap protected lemma inducing : Inducing f := (hf.homeomorph f).inducing -protected lemma quotientMap : QuotientMap f := (hf.homeomorph f).quotientMap +lemma isQuotientMap : IsQuotientMap f := (hf.homeomorph f).isQuotientMap protected lemma embedding : Embedding f := (hf.homeomorph f).embedding lemma isOpenEmbedding : IsOpenEmbedding f := (hf.homeomorph f).isOpenEmbedding lemma isClosedEmbedding : IsClosedEmbedding f := (hf.homeomorph f).isClosedEmbedding lemma isDenseEmbedding : IsDenseEmbedding f := (hf.homeomorph f).isDenseEmbedding +@[deprecated (since := "2024-10-22")] +alias quotientMap := isQuotientMap + @[deprecated (since := "2024-10-20")] alias closedEmbedding := isClosedEmbedding @[deprecated (since := "2024-10-18")] alias openEmbedding := isOpenEmbedding diff --git a/Mathlib/Topology/Inseparable.lean b/Mathlib/Topology/Inseparable.lean index 919eeb7228337..99dce9c03e378 100644 --- a/Mathlib/Topology/Inseparable.lean +++ b/Mathlib/Topology/Inseparable.lean @@ -515,8 +515,11 @@ namespace SeparationQuotient /-- The natural map from a topological space to its separation quotient. -/ def mk : X → SeparationQuotient X := Quotient.mk'' -theorem quotientMap_mk : QuotientMap (mk : X → SeparationQuotient X) := - quotientMap_quot_mk +theorem isQuotientMap_mk : IsQuotientMap (mk : X → SeparationQuotient X) := + isQuotientMap_quot_mk + +@[deprecated (since := "2024-10-22")] +alias quotientMap_mk := isQuotientMap_mk @[fun_prop, continuity] theorem continuous_mk : Continuous (mk : X → SeparationQuotient X) := @@ -552,7 +555,7 @@ theorem preimage_image_mk_open (hs : IsOpen s) : mk ⁻¹' (mk '' s) = s := by exact ((mk_eq_mk.1 hxy).mem_open_iff hs).1 hys theorem isOpenMap_mk : IsOpenMap (mk : X → SeparationQuotient X) := fun s hs => - quotientMap_mk.isOpen_preimage.1 <| by rwa [preimage_image_mk_open hs] + isQuotientMap_mk.isOpen_preimage.1 <| by rwa [preimage_image_mk_open hs] theorem isOpenQuotientMap_mk : IsOpenQuotientMap (mk : X → SeparationQuotient X) := ⟨surjective_mk, continuous_mk, isOpenMap_mk⟩ @@ -608,8 +611,11 @@ 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 → _) := - (isOpenQuotientMap_mk.prodMap isOpenQuotientMap_mk).quotientMap +theorem isQuotientMap_prodMap_mk : IsQuotientMap (Prod.map mk mk : X × Y → _) := + (isOpenQuotientMap_mk.prodMap isOpenQuotientMap_mk).isQuotientMap + +@[deprecated (since := "2024-10-22")] +alias quotientMap_prodMap_mk := isQuotientMap_prodMap_mk /-- 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/AddCircle.lean b/Mathlib/Topology/Instances/AddCircle.lean index 969b2df626d4d..4c5e3613f132b 100644 --- a/Mathlib/Topology/Instances/AddCircle.lean +++ b/Mathlib/Topology/Instances/AddCircle.lean @@ -591,7 +591,7 @@ homeomorphism of topological spaces. -/ def homeoIccQuot [TopologicalSpace 𝕜] [OrderTopology 𝕜] : 𝕋 ≃ₜ Quot (EndpointIdent p a) where toEquiv := equivIccQuot p a continuous_toFun := by - simp_rw [quotientMap_quotient_mk'.continuous_iff, continuous_iff_continuousAt, + simp_rw [isQuotientMap_quotient_mk'.continuous_iff, continuous_iff_continuousAt, continuousAt_iff_continuous_left_right] intro x; constructor on_goal 1 => erw [equivIccQuot_comp_mk_eq_toIocMod] diff --git a/Mathlib/Topology/Maps/Basic.lean b/Mathlib/Topology/Maps/Basic.lean index 7933f9ee54473..16c18a15ebc2f 100644 --- a/Mathlib/Topology/Maps/Basic.lean +++ b/Mathlib/Topology/Maps/Basic.lean @@ -25,7 +25,7 @@ This file introduces the following properties of a map `f : X → Y` between top * `IsClosedEmbedding f` similarly means `f` is an embedding with closed image, so it identifies `X` with a closed subspace of `Y`. Equivalently, `f` is an embedding and a closed map. -* `QuotientMap f` is the dual condition to `Embedding f`: `f` is surjective and the topology +* `IsQuotientMap f` is the dual condition to `Embedding f`: `f` is surjective and the topology on `Y` is the one coinduced via `f` from the topology on `X`. Equivalently, `f` identifies `Y` with a quotient of `X`. Quotient maps are also sometimes known as identification maps. @@ -228,53 +228,61 @@ theorem Embedding.of_subsingleton [Subsingleton X] (f : X → Y) : Embedding f : end Embedding -section QuotientMap +section IsQuotientMap variable [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] -theorem quotientMap_iff : QuotientMap f ↔ Surjective f ∧ ∀ s : Set Y, IsOpen s ↔ IsOpen (f ⁻¹' s) := - (quotientMap_iff' _).trans <| and_congr Iff.rfl TopologicalSpace.ext_iff +lemma isQuotientMap_iff : IsQuotientMap f ↔ Surjective f ∧ ∀ s, IsOpen s ↔ IsOpen (f ⁻¹' s) := + (isQuotientMap_iff' _).trans <| and_congr Iff.rfl TopologicalSpace.ext_iff -theorem quotientMap_iff_closed : - QuotientMap f ↔ Surjective f ∧ ∀ s : Set Y, IsClosed s ↔ IsClosed (f ⁻¹' s) := - quotientMap_iff.trans <| Iff.rfl.and <| compl_surjective.forall.trans <| by +@[deprecated (since := "2024-10-22")] +alias quotientMap_iff := isQuotientMap_iff + +theorem isQuotientMap_iff_closed : + IsQuotientMap f ↔ Surjective f ∧ ∀ s : Set Y, IsClosed s ↔ IsClosed (f ⁻¹' s) := + isQuotientMap_iff.trans <| Iff.rfl.and <| compl_surjective.forall.trans <| by simp only [isOpen_compl_iff, preimage_compl] -namespace QuotientMap +@[deprecated (since := "2024-10-22")] +alias quotientMap_iff_closed := isQuotientMap_iff_closed + +namespace IsQuotientMap -protected theorem id : QuotientMap (@id X) := +protected theorem id : IsQuotientMap (@id X) := ⟨fun x => ⟨x, rfl⟩, coinduced_id.symm⟩ -protected theorem comp (hg : QuotientMap g) (hf : QuotientMap f) : QuotientMap (g ∘ f) := +protected theorem comp (hg : IsQuotientMap g) (hf : IsQuotientMap f) : IsQuotientMap (g ∘ f) := ⟨hg.surjective.comp hf.surjective, by rw [hg.eq_coinduced, hf.eq_coinduced, coinduced_compose]⟩ -protected theorem of_quotientMap_compose (hf : Continuous f) (hg : Continuous g) - (hgf : QuotientMap (g ∘ f)) : QuotientMap g := +protected theorem of_comp (hf : Continuous f) (hg : Continuous g) + (hgf : IsQuotientMap (g ∘ f)) : IsQuotientMap g := ⟨hgf.1.of_comp, le_antisymm (by rw [hgf.eq_coinduced, ← coinduced_compose]; exact coinduced_mono hf.coinduced_le) hg.coinduced_le⟩ +@[deprecated (since := "2024-10-22")] +alias of_quotientMap_compose := IsQuotientMap.of_comp + theorem of_inverse {g : Y → X} (hf : Continuous f) (hg : Continuous g) (h : LeftInverse g f) : - QuotientMap g := - QuotientMap.of_quotientMap_compose hf hg <| h.comp_eq_id.symm ▸ QuotientMap.id + IsQuotientMap g := .of_comp hf hg <| h.comp_eq_id.symm ▸ IsQuotientMap.id -protected theorem continuous_iff (hf : QuotientMap f) : Continuous g ↔ Continuous (g ∘ f) := by +protected theorem continuous_iff (hf : IsQuotientMap f) : Continuous g ↔ Continuous (g ∘ f) := by rw [continuous_iff_coinduced_le, continuous_iff_coinduced_le, hf.eq_coinduced, coinduced_compose] -protected theorem continuous (hf : QuotientMap f) : Continuous f := +protected theorem continuous (hf : IsQuotientMap f) : Continuous f := hf.continuous_iff.mp continuous_id -protected theorem isOpen_preimage (hf : QuotientMap f) {s : Set Y} : IsOpen (f ⁻¹' s) ↔ IsOpen s := - ((quotientMap_iff.1 hf).2 s).symm +protected lemma isOpen_preimage (hf : IsQuotientMap f) {s : Set Y} : IsOpen (f ⁻¹' s) ↔ IsOpen s := + ((isQuotientMap_iff.1 hf).2 s).symm -protected theorem isClosed_preimage (hf : QuotientMap f) {s : Set Y} : +protected theorem isClosed_preimage (hf : IsQuotientMap f) {s : Set Y} : IsClosed (f ⁻¹' s) ↔ IsClosed s := - ((quotientMap_iff_closed.1 hf).2 s).symm + ((isQuotientMap_iff_closed.1 hf).2 s).symm -end QuotientMap +end IsQuotientMap -end QuotientMap +end IsQuotientMap section OpenMap variable [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] @@ -326,11 +334,14 @@ theorem of_inverse {f' : Y → X} (h : Continuous f') (l_inv : LeftInverse f f') of_sections fun _ => ⟨f', h.continuousAt, r_inv _, l_inv⟩ /-- A continuous surjective open map is a quotient map. -/ -theorem to_quotientMap (open_map : IsOpenMap f) (cont : Continuous f) (surj : Surjective f) : - QuotientMap f := - quotientMap_iff.2 +theorem isQuotientMap (open_map : IsOpenMap f) (cont : Continuous f) (surj : Surjective f) : + IsQuotientMap f := + isQuotientMap_iff.2 ⟨surj, fun s => ⟨fun h => h.preimage cont, fun h => surj.image_preimage s ▸ open_map _ h⟩⟩ +@[deprecated (since := "2024-10-22")] +alias to_quotientMap := isQuotientMap + theorem interior_preimage_subset_preimage_interior (hf : IsOpenMap f) {s : Set Y} : interior (f ⁻¹' s) ⊆ f ⁻¹' interior s := hf.mapsTo_interior (mapsTo_preimage _ _) @@ -419,11 +430,14 @@ theorem isClosed_range (hf : IsClosedMap f) : IsClosed (range f) := @[deprecated (since := "2024-03-17")] alias closed_range := isClosed_range -theorem to_quotientMap (hcl : IsClosedMap f) (hcont : Continuous f) - (hsurj : Surjective f) : QuotientMap f := - quotientMap_iff_closed.2 ⟨hsurj, fun s => +theorem isQuotientMap (hcl : IsClosedMap f) (hcont : Continuous f) + (hsurj : Surjective f) : IsQuotientMap f := + isQuotientMap_iff_closed.2 ⟨hsurj, fun s => ⟨fun hs => hs.preimage hcont, fun hs => hsurj.image_preimage s ▸ hcl _ hs⟩⟩ +@[deprecated (since := "2024-10-22")] +alias to_quotientMap := isQuotientMap + end IsClosedMap theorem Inducing.isClosedMap (hf : Inducing f) (h : IsClosed (range f)) : IsClosedMap f := by diff --git a/Mathlib/Topology/Maps/OpenQuotient.lean b/Mathlib/Topology/Maps/OpenQuotient.lean index 88433f7a190fd..4f593b59638a1 100644 --- a/Mathlib/Topology/Maps/OpenQuotient.lean +++ b/Mathlib/Topology/Maps/OpenQuotient.lean @@ -32,15 +32,24 @@ namespace IsOpenQuotientMap protected theorem id : IsOpenQuotientMap (id : X → X) := ⟨surjective_id, continuous_id, .id⟩ /-- An open quotient map is a quotient map. -/ -theorem quotientMap (h : IsOpenQuotientMap f) : QuotientMap f := - h.isOpenMap.to_quotientMap h.continuous h.surjective +theorem isQuotientMap (h : IsOpenQuotientMap f) : IsQuotientMap f := + h.isOpenMap.isQuotientMap h.continuous h.surjective -theorem iff_isOpenMap_quotientMap : IsOpenQuotientMap f ↔ IsOpenMap f ∧ QuotientMap f := - ⟨fun h ↦ ⟨h.isOpenMap, h.quotientMap⟩, fun ⟨ho, hq⟩ ↦ ⟨hq.surjective, hq.continuous, ho⟩⟩ +@[deprecated (since := "2024-10-22")] +alias quotientMap := isQuotientMap -theorem of_isOpenMap_quotientMap (ho : IsOpenMap f) (hq : QuotientMap f) : +theorem iff_isOpenMap_isQuotientMap : IsOpenQuotientMap f ↔ IsOpenMap f ∧ IsQuotientMap f := + ⟨fun h ↦ ⟨h.isOpenMap, h.isQuotientMap⟩, fun ⟨ho, hq⟩ ↦ ⟨hq.surjective, hq.continuous, ho⟩⟩ + +@[deprecated (since := "2024-10-22")] +alias iff_isOpenMap_quotientMap := iff_isOpenMap_isQuotientMap + +theorem of_isOpenMap_isQuotientMap (ho : IsOpenMap f) (hq : IsQuotientMap f) : IsOpenQuotientMap f := - iff_isOpenMap_quotientMap.2 ⟨ho, hq⟩ + iff_isOpenMap_isQuotientMap.2 ⟨ho, hq⟩ + +@[deprecated (since := "2024-10-22")] +alias of_isOpenMap_quotientMap := of_isOpenMap_isQuotientMap theorem comp {g : Y → Z} (hg : IsOpenQuotientMap g) (hf : IsOpenQuotientMap f) : IsOpenQuotientMap (g ∘ f) := @@ -51,7 +60,7 @@ theorem map_nhds_eq (h : IsOpenQuotientMap f) (x : X) : map f (𝓝 x) = 𝓝 (f theorem continuous_comp_iff (h : IsOpenQuotientMap f) {g : Y → Z} : Continuous (g ∘ f) ↔ Continuous g := - h.quotientMap.continuous_iff.symm + h.isQuotientMap.continuous_iff.symm theorem continuousAt_comp_iff (h : IsOpenQuotientMap f) {g : Y → Z} {x : X} : ContinuousAt (g ∘ f) x ↔ ContinuousAt g (f x) := by diff --git a/Mathlib/Topology/Order/ProjIcc.lean b/Mathlib/Topology/Order/ProjIcc.lean index 993c017bb989f..7c90ab75b6439 100644 --- a/Mathlib/Topology/Order/ProjIcc.lean +++ b/Mathlib/Topology/Order/ProjIcc.lean @@ -29,13 +29,16 @@ variable [TopologicalSpace α] [OrderTopology α] [TopologicalSpace β] [Topolog theorem continuous_projIcc : Continuous (projIcc a b h) := (continuous_const.max <| continuous_const.min continuous_id).subtype_mk _ -theorem quotientMap_projIcc : QuotientMap (projIcc a b h) := - quotientMap_iff.2 ⟨projIcc_surjective h, fun s => +theorem isQuotientMap_projIcc : IsQuotientMap (projIcc a b h) := + isQuotientMap_iff.2 ⟨projIcc_surjective h, fun s => ⟨fun hs => hs.preimage continuous_projIcc, fun hs => ⟨_, hs, by ext; simp⟩⟩⟩ +@[deprecated (since := "2024-10-22")] +alias quotientMap_projIcc := isQuotientMap_projIcc + @[simp] theorem continuous_IccExtend_iff {f : Icc a b → β} : Continuous (IccExtend h f) ↔ Continuous f := - quotientMap_projIcc.continuous_iff.symm + isQuotientMap_projIcc.continuous_iff.symm /-- See Note [continuity lemma statement]. -/ protected theorem Continuous.IccExtend {f : γ → Icc a b → β} {g : γ → α} (hf : Continuous ↿f) diff --git a/Mathlib/Topology/Separation.lean b/Mathlib/Topology/Separation.lean index 386aaf8de79c8..a99e1b615e460 100644 --- a/Mathlib/Topology/Separation.lean +++ b/Mathlib/Topology/Separation.lean @@ -1864,9 +1864,12 @@ theorem Continuous.isClosedEmbedding [CompactSpace X] [T2Space Y] {f : X → Y} alias Continuous.closedEmbedding := Continuous.isClosedEmbedding /-- A continuous surjective map from a compact space to a Hausdorff space is a quotient map. -/ -theorem QuotientMap.of_surjective_continuous [CompactSpace X] [T2Space Y] {f : X → Y} - (hsurj : Surjective f) (hcont : Continuous f) : QuotientMap f := - hcont.isClosedMap.to_quotientMap hcont hsurj +theorem IsQuotientMap.of_surjective_continuous [CompactSpace X] [T2Space Y] {f : X → Y} + (hsurj : Surjective f) (hcont : Continuous f) : IsQuotientMap f := + hcont.isClosedMap.isQuotientMap hcont hsurj + +@[deprecated (since := "2024-10-22")] +alias QuotientMap.of_surjective_continuous := IsQuotientMap.of_surjective_continuous theorem isPreirreducible_iff_subsingleton [T2Space X] {S : Set X} : IsPreirreducible S ↔ S.Subsingleton := by @@ -2659,7 +2662,7 @@ instance ConnectedComponents.t2 [T2Space X] [CompactSpace X] : T2Space (Connecte have hU : IsClopen U := isClopen_biInter_finset fun i _ => i.2.1 exact ⟨U, (↑) '' U, hU, ha, subset_iInter₂ fun s _ => s.2.1.connectedComponent_subset s.2.2, (connectedComponents_preimage_image U).symm ▸ hU.biUnion_connectedComponent_eq⟩ - rw [ConnectedComponents.quotientMap_coe.isClopen_preimage] at hU + rw [ConnectedComponents.isQuotientMap_coe.isClopen_preimage] at hU refine ⟨Vᶜ, V, hU.compl.isOpen, hU.isOpen, ?_, hb mem_connectedComponent, disjoint_compl_left⟩ exact fun h => flip Set.Nonempty.ne_empty ha ⟨a, mem_connectedComponent, h⟩ diff --git a/Mathlib/Topology/Sequences.lean b/Mathlib/Topology/Sequences.lean index f19fd89e1847d..d1805b3da9465 100644 --- a/Mathlib/Topology/Sequences.lean +++ b/Mathlib/Topology/Sequences.lean @@ -207,14 +207,17 @@ protected theorem SequentialSpace.sup {X} {t₁ t₂ : TopologicalSpace X} rw [sup_eq_iSup] exact .iSup <| Bool.forall_bool.2 ⟨h₂, h₁⟩ -theorem QuotientMap.sequentialSpace [SequentialSpace X] {f : X → Y} (hf : QuotientMap f) : +theorem IsQuotientMap.sequentialSpace [SequentialSpace X] {f : X → Y} (hf : IsQuotientMap f) : SequentialSpace Y := hf.2.symm ▸ .coinduced f +@[deprecated (since := "2024-10-22")] +alias QuotientMap.sequentialSpace := IsQuotientMap.sequentialSpace + /-- The quotient of a sequential space is a sequential space. -/ instance Quotient.instSequentialSpace [SequentialSpace X] {s : Setoid X} : SequentialSpace (Quotient s) := - quotientMap_quot_mk.sequentialSpace + isQuotientMap_quot_mk.sequentialSpace /-- The sum (disjoint union) of two sequential spaces is a sequential space. -/ instance Sum.instSequentialSpace [SequentialSpace X] [SequentialSpace Y] : diff --git a/scripts/no_lints_prime_decls.txt b/scripts/no_lints_prime_decls.txt index b232c847b8a7a..ddd9bf20bcedc 100644 --- a/scripts/no_lints_prime_decls.txt +++ b/scripts/no_lints_prime_decls.txt @@ -4039,7 +4039,7 @@ Quotient.liftOn₂'_mk'' Quotient.liftOn'_mk'' Quotient.map₂'_mk'' Quotient.map'_mk'' -quotientMap_quotient_mk' +isQuotientMap_quotient_mk' Quotient.mk_out' Quotient.out_eq' Quotient.sound' From 77ffa0d4546c2cb478659b628728ec4a58e6b669 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Wed, 23 Oct 2024 04:26:17 +0000 Subject: [PATCH 317/505] chore(SetTheory/Cardinal/Basic): deprecate duplicate `Nat.cast` lemmas (#17850) Spotted thanks to #17827. --- Mathlib/Data/Finite/Card.lean | 2 +- Mathlib/FieldTheory/Fixed.lean | 4 +- Mathlib/LinearAlgebra/Dimension/Finite.lean | 4 +- .../Dimension/FreeAndStrongRankCondition.lean | 6 +-- .../Dimension/StrongRankCondition.lean | 6 +-- .../LinearAlgebra/FiniteDimensional/Defs.lean | 2 +- Mathlib/NumberTheory/RamificationInertia.lean | 2 +- Mathlib/SetTheory/Cardinal/Basic.lean | 50 +++++++++---------- Mathlib/SetTheory/Ordinal/Basic.lean | 2 +- .../Topology/MetricSpace/GromovHausdorff.lean | 2 +- 10 files changed, 40 insertions(+), 40 deletions(-) diff --git a/Mathlib/Data/Finite/Card.lean b/Mathlib/Data/Finite/Card.lean index 397130772c82f..fb6afe0f1c2f9 100644 --- a/Mathlib/Data/Finite/Card.lean +++ b/Mathlib/Data/Finite/Card.lean @@ -181,7 +181,7 @@ theorem card_union_le (s t : Set α) : Nat.card (↥(s ∪ t)) ≤ Nat.card s + cases' _root_.finite_or_infinite (↥(s ∪ t)) with h h · rw [finite_coe_iff, finite_union, ← finite_coe_iff, ← finite_coe_iff] at h cases h - rw [← Cardinal.natCast_le, Nat.cast_add, Finite.cast_card_eq_mk, Finite.cast_card_eq_mk, + rw [← @Nat.cast_le Cardinal, Nat.cast_add, Finite.cast_card_eq_mk, Finite.cast_card_eq_mk, Finite.cast_card_eq_mk] exact Cardinal.mk_union_le s t · exact Nat.card_eq_zero_of_infinite.trans_le (zero_le _) diff --git a/Mathlib/FieldTheory/Fixed.lean b/Mathlib/FieldTheory/Fixed.lean index 4b86ac581b736..92f028d565270 100644 --- a/Mathlib/FieldTheory/Fixed.lean +++ b/Mathlib/FieldTheory/Fixed.lean @@ -246,7 +246,7 @@ theorem minpoly_eq_minpoly : minpoly G F x = _root_.minpoly (FixedPoints.subfiel theorem rank_le_card : Module.rank (FixedPoints.subfield G F) F ≤ Fintype.card G := rank_le fun s hs => by simpa only [rank_fun', Cardinal.mk_coe_finset, Finset.coe_sort_coe, Cardinal.lift_natCast, - Cardinal.natCast_le] using + Nat.cast_le] using (linearIndependent_smul_of_linearIndependent G F hs).cardinal_lift_le_rank end Fintype @@ -282,7 +282,7 @@ instance : FiniteDimensional (subfield G F) F := by end Finite theorem finrank_le_card [Fintype G] : finrank (subfield G F) F ≤ Fintype.card G := by - rw [← Cardinal.natCast_le, finrank_eq_rank] + rw [← @Nat.cast_le Cardinal, finrank_eq_rank] apply rank_le_card end FixedPoints diff --git a/Mathlib/LinearAlgebra/Dimension/Finite.lean b/Mathlib/LinearAlgebra/Dimension/Finite.lean index bf110e9d64974..8cce5dddc805b 100644 --- a/Mathlib/LinearAlgebra/Dimension/Finite.lean +++ b/Mathlib/LinearAlgebra/Dimension/Finite.lean @@ -233,7 +233,7 @@ lemma exists_finset_linearIndependent_of_le_finrank {n : ℕ} (hn : n ≤ finran · rw [le_zero_iff.mp (hn.trans_eq h)] exact ⟨∅, rfl, by convert linearIndependent_empty R M using 2 <;> aesop⟩ exact exists_finset_linearIndependent_of_le_rank - ((natCast_le.mpr hn).trans_eq (cast_toNat_of_lt_aleph0 (toNat_ne_zero.mp h).2)) + ((Nat.cast_le.mpr hn).trans_eq (cast_toNat_of_lt_aleph0 (toNat_ne_zero.mp h).2)) lemma exists_linearIndependent_of_le_finrank {n : ℕ} (hn : n ≤ finrank R M) : ∃ f : Fin n → M, LinearIndependent R f := @@ -463,7 +463,7 @@ theorem Submodule.rank_eq_zero [Nontrivial R] [NoZeroSMulDivisors R M] {S : Subm theorem Submodule.finrank_eq_zero [StrongRankCondition R] [NoZeroSMulDivisors R M] {S : Submodule R M} [Module.Finite R S] : finrank R S = 0 ↔ S = ⊥ := by - rw [← Submodule.rank_eq_zero, ← finrank_eq_rank, ← @Nat.cast_zero Cardinal, Cardinal.natCast_inj] + rw [← Submodule.rank_eq_zero, ← finrank_eq_rank, ← @Nat.cast_zero Cardinal, Nat.cast_inj] @[simp] lemma Submodule.one_le_finrank_iff [StrongRankCondition R] [NoZeroSMulDivisors R M] diff --git a/Mathlib/LinearAlgebra/Dimension/FreeAndStrongRankCondition.lean b/Mathlib/LinearAlgebra/Dimension/FreeAndStrongRankCondition.lean index 17254b66084e3..43e1a1eedf450 100644 --- a/Mathlib/LinearAlgebra/Dimension/FreeAndStrongRankCondition.lean +++ b/Mathlib/LinearAlgebra/Dimension/FreeAndStrongRankCondition.lean @@ -202,16 +202,16 @@ there is some `v : V` so every vector is a multiple of `v`. -/ theorem finrank_le_one_iff [Module.Free K V] [Module.Finite K V] : finrank K V ≤ 1 ↔ ∃ v : V, ∀ w : V, ∃ c : K, c • v = w := by - rw [← rank_le_one_iff, ← finrank_eq_rank, ← natCast_le, Nat.cast_one] + rw [← rank_le_one_iff, ← finrank_eq_rank, Nat.cast_le_one] theorem Submodule.finrank_le_one_iff_isPrincipal (W : Submodule K V) [Module.Free K W] [Module.Finite K W] : finrank K W ≤ 1 ↔ W.IsPrincipal := by - rw [← W.rank_le_one_iff_isPrincipal, ← finrank_eq_rank, ← natCast_le, Nat.cast_one] + rw [← W.rank_le_one_iff_isPrincipal, ← finrank_eq_rank, Nat.cast_le_one] theorem Module.finrank_le_one_iff_top_isPrincipal [Module.Free K V] [Module.Finite K V] : finrank K V ≤ 1 ↔ (⊤ : Submodule K V).IsPrincipal := by - rw [← Module.rank_le_one_iff_top_isPrincipal, ← finrank_eq_rank, ← natCast_le, Nat.cast_one] + rw [← Module.rank_le_one_iff_top_isPrincipal, ← finrank_eq_rank, Nat.cast_le_one] variable (K V) in theorem lift_cardinal_mk_eq_lift_cardinal_mk_field_pow_lift_rank [Module.Free K V] diff --git a/Mathlib/LinearAlgebra/Dimension/StrongRankCondition.lean b/Mathlib/LinearAlgebra/Dimension/StrongRankCondition.lean index 7b753996882f1..a84c1fb77e7b2 100644 --- a/Mathlib/LinearAlgebra/Dimension/StrongRankCondition.lean +++ b/Mathlib/LinearAlgebra/Dimension/StrongRankCondition.lean @@ -64,7 +64,7 @@ theorem mk_eq_mk_of_basis (v : Basis ι R M) (v' : Basis ι' R M) : cases nonempty_fintype ι' -- We clean up a little: rw [Cardinal.mk_fintype, Cardinal.mk_fintype] - simp only [Cardinal.lift_natCast, Cardinal.natCast_inj] + simp only [Cardinal.lift_natCast, Nat.cast_inj] -- Now we can use invariant basis number to show they have the same cardinality. apply card_eq_of_linearEquiv R exact @@ -122,7 +122,7 @@ theorem basis_le_span' {ι : Type*} (b : Basis ι R M) {w : Set M} [Fintype w] ( haveI := basis_finite_of_finite_spans w (toFinite _) s b cases nonempty_fintype ι rw [Cardinal.mk_fintype ι] - simp only [Cardinal.natCast_le] + simp only [Nat.cast_le] exact Basis.le_span'' b s -- Note that if `R` satisfies the strong rank condition, @@ -207,7 +207,7 @@ theorem linearIndependent_le_span' {ι : Type*} (v : ι → M) (i : LinearIndepe haveI : Finite ι := i.finite_of_le_span_finite v w s letI := Fintype.ofFinite ι rw [Cardinal.mk_fintype] - simp only [Cardinal.natCast_le] + simp only [Nat.cast_le] exact linearIndependent_le_span_aux' v i w s /-- If `R` satisfies the strong rank condition, diff --git a/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean b/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean index 372c3c0e74552..bcbbfded66f88 100644 --- a/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean +++ b/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean @@ -488,7 +488,7 @@ variable [DivisionRing K] [AddCommGroup V] [Module K V] {V₂ : Type v'} [AddCom theorem surjective_of_injective [FiniteDimensional K V] {f : V →ₗ[K] V} (hinj : Injective f) : Surjective f := by have h := rank_range_of_injective _ hinj - rw [← finrank_eq_rank, ← finrank_eq_rank, natCast_inj] at h + rw [← finrank_eq_rank, ← finrank_eq_rank, Nat.cast_inj] at h exact range_eq_top.1 (eq_top_of_finrank_eq h) /-- The image under an onto linear map of a finite-dimensional space is also finite-dimensional. -/ diff --git a/Mathlib/NumberTheory/RamificationInertia.lean b/Mathlib/NumberTheory/RamificationInertia.lean index 10d1ba5cf00ef..a69dfce45bd5a 100644 --- a/Mathlib/NumberTheory/RamificationInertia.lean +++ b/Mathlib/NumberTheory/RamificationInertia.lean @@ -639,7 +639,7 @@ theorem finrank_prime_pow_ramificationIdx [IsDedekindDomain S] (hP0 : P ≠ ⊥) by_cases hP : FiniteDimensional (R ⧸ p) (S ⧸ P) · haveI := hP haveI := (finiteDimensional_iff_of_rank_eq_nsmul he hdim).mpr hP - refine Cardinal.natCast_injective ?_ + apply @Nat.cast_injective Cardinal rw [finrank_eq_rank', Nat.cast_mul, finrank_eq_rank', hdim, nsmul_eq_mul] have hPe := mt (finiteDimensional_iff_of_rank_eq_nsmul he hdim).mp hP simp only [finrank_of_infinite_dimensional hP, finrank_of_infinite_dimensional hPe, diff --git a/Mathlib/SetTheory/Cardinal/Basic.lean b/Mathlib/SetTheory/Cardinal/Basic.lean index 006f281adb63c..10ebaa481ad46 100644 --- a/Mathlib/SetTheory/Cardinal/Basic.lean +++ b/Mathlib/SetTheory/Cardinal/Basic.lean @@ -563,10 +563,13 @@ theorem power_mul {a b c : Cardinal} : a ^ (b * c) = (a ^ b) ^ c := by rw [mul_comm b c] exact inductionOn₃ a b c fun α β γ => mk_congr <| Equiv.curry γ β α -@[simp] -theorem pow_cast_right (a : Cardinal.{u}) (n : ℕ) : a ^ (↑n : Cardinal.{u}) = a ^ n := +@[simp, norm_cast] +theorem power_natCast (a : Cardinal.{u}) (n : ℕ) : a ^ (↑n : Cardinal.{u}) = a ^ n := rfl +@[deprecated (since := "2024-10-16")] +alias power_cast_right := power_natCast + @[simp] theorem lift_one : lift 1 = 1 := mk_eq_one _ @@ -1381,36 +1384,33 @@ theorem mk_finset_of_fintype [Fintype α] : #(Finset α) = 2 ^ Fintype.card α : theorem card_le_of_finset {α} (s : Finset α) : (s.card : Cardinal) ≤ #α := @mk_coe_finset _ s ▸ mk_set_le _ -@[norm_cast] -theorem natCast_pow {m n : ℕ} : (↑(m ^ n) : Cardinal) = (↑m : Cardinal) ^ (↑n : Cardinal) := by - induction n <;> simp [pow_succ, power_add, *, Pow.pow] +instance : CharZero Cardinal := by + refine ⟨fun a b h ↦ ?_⟩ + rwa [← lift_mk_fin, ← lift_mk_fin, lift_inj, Cardinal.eq, ← Fintype.card_eq, + Fintype.card_fin, Fintype.card_fin] at h -@[norm_cast] -theorem natCast_le {m n : ℕ} : (m : Cardinal) ≤ n ↔ m ≤ n := by - rw [← lift_mk_fin, ← lift_mk_fin, lift_le, le_def, Function.Embedding.nonempty_iff_card_le, - Fintype.card_fin, Fintype.card_fin] +@[deprecated Nat.cast_le (since := "2024-10-16")] +theorem natCast_le {m n : ℕ} : (m : Cardinal) ≤ n ↔ m ≤ n := Nat.cast_le -@[norm_cast] -theorem natCast_lt {m n : ℕ} : (m : Cardinal) < n ↔ m < n := by - rw [lt_iff_le_not_le, ← not_le] - simp only [natCast_le, not_le, and_iff_right_iff_imp] - exact fun h ↦ le_of_lt h +@[deprecated Nat.cast_lt (since := "2024-10-16")] +theorem natCast_lt {m n : ℕ} : (m : Cardinal) < n ↔ m < n := Nat.cast_lt -instance : CharZero Cardinal := - ⟨StrictMono.injective fun _ _ => natCast_lt.2⟩ +@[deprecated Nat.cast_inj (since := "2024-10-16")] +theorem natCast_inj {m n : ℕ} : (m : Cardinal) = n ↔ m = n := Nat.cast_inj -theorem natCast_inj {m n : ℕ} : (m : Cardinal) = n ↔ m = n := - Nat.cast_inj +@[deprecated Nat.cast_injective (since := "2024-10-16")] +theorem natCast_injective : Injective ((↑) : ℕ → Cardinal) := Nat.cast_injective -theorem natCast_injective : Injective ((↑) : ℕ → Cardinal) := - Nat.cast_injective +@[deprecated Nat.cast_pow (since := "2024-10-16")] +theorem natCast_pow {m n : ℕ} : (↑(m ^ n) : Cardinal) = (↑m : Cardinal) ^ (↑n : Cardinal) := + Nat.cast_pow m n @[norm_cast] theorem nat_succ (n : ℕ) : (n.succ : Cardinal) = succ ↑n := by rw [Nat.cast_succ] refine (add_one_le_succ _).antisymm (succ_le_of_lt ?_) rw [← Nat.cast_succ] - exact natCast_lt.2 (Nat.lt_succ_self _) + exact Nat.cast_lt.2 (Nat.lt_succ_self _) lemma succ_natCast (n : ℕ) : Order.succ (n : Cardinal) = n + 1 := by rw [← Cardinal.nat_succ] @@ -1490,7 +1490,7 @@ theorem aleph0_le {c : Cardinal} : ℵ₀ ≤ c ↔ ∀ n : ℕ, ↑n ≤ c := ⟨fun h _ => (nat_lt_aleph0 _).le.trans h, fun h => le_of_not_lt fun hn => by rcases lt_aleph0.1 hn with ⟨n, rfl⟩ - exact (Nat.lt_succ_self _).not_le (natCast_le.1 (h (n + 1)))⟩ + exact (Nat.lt_succ_self _).not_le (Nat.cast_le.1 (h (n + 1)))⟩ theorem isSuccPrelimit_aleph0 : IsSuccPrelimit ℵ₀ := isSuccPrelimit_of_succ_lt fun a ha => by @@ -1646,7 +1646,7 @@ theorem mul_lt_aleph0_iff_of_ne_zero {a b : Cardinal} (ha : a ≠ 0) (hb : b ≠ theorem power_lt_aleph0 {a b : Cardinal} (ha : a < ℵ₀) (hb : b < ℵ₀) : a ^ b < ℵ₀ := match a, b, lt_aleph0.1 ha, lt_aleph0.1 hb with - | _, _, ⟨m, rfl⟩, ⟨n, rfl⟩ => by rw [← natCast_pow]; apply nat_lt_aleph0 + | _, _, ⟨m, rfl⟩, ⟨n, rfl⟩ => by rw [power_natCast, ← Nat.cast_pow]; apply nat_lt_aleph0 theorem eq_one_iff_unique {α : Type*} : #α = 1 ↔ Subsingleton α ∧ Nonempty α := calc @@ -1691,7 +1691,7 @@ theorem aleph0_mul_aleph0 : ℵ₀ * ℵ₀ = ℵ₀ := theorem nat_mul_aleph0 {n : ℕ} (hn : n ≠ 0) : ↑n * ℵ₀ = ℵ₀ := le_antisymm (lift_mk_fin n ▸ mk_le_aleph0) <| le_mul_of_one_le_left (zero_le _) <| by - rwa [← Nat.cast_one, natCast_le, Nat.one_le_iff_ne_zero] + rwa [← Nat.cast_one, Nat.cast_le, Nat.one_le_iff_ne_zero] @[simp] theorem aleph0_mul_nat {n : ℕ} (hn : n ≠ 0) : ℵ₀ * n = ℵ₀ := by rw [mul_comm, nat_mul_aleph0 hn] @@ -2072,7 +2072,7 @@ theorem exists_not_mem_of_length_lt {α : Type*} (l : List α) (h : ↑l.length #α = #(Set.univ : Set α) := mk_univ.symm _ ≤ #l.toFinset := mk_le_mk_of_subset fun x _ => List.mem_toFinset.mpr (h x) _ = l.toFinset.card := Cardinal.mk_coe_finset - _ ≤ l.length := Cardinal.natCast_le.mpr (List.toFinset_card_le l) + _ ≤ l.length := Nat.cast_le.mpr (List.toFinset_card_le l) theorem three_le {α : Type*} (h : 3 ≤ #α) (x : α) (y : α) : ∃ z : α, z ≠ x ∧ z ≠ y := by have : ↑(3 : ℕ) ≤ #α := by simpa using h diff --git a/Mathlib/SetTheory/Ordinal/Basic.lean b/Mathlib/SetTheory/Ordinal/Basic.lean index 5c5d1915ecefc..9ccd58c41f693 100644 --- a/Mathlib/SetTheory/Ordinal/Basic.lean +++ b/Mathlib/SetTheory/Ordinal/Basic.lean @@ -1196,7 +1196,7 @@ theorem ord_nat (n : ℕ) : ord n = n := (by induction' n with n IH · apply Ordinal.zero_le - · exact succ_le_of_lt (IH.trans_lt <| ord_lt_ord.2 <| natCast_lt.2 (Nat.lt_succ_self n))) + · exact succ_le_of_lt (IH.trans_lt <| ord_lt_ord.2 <| Nat.cast_lt.2 (Nat.lt_succ_self n))) @[simp] theorem ord_one : ord 1 = 1 := by simpa using ord_nat 1 diff --git a/Mathlib/Topology/MetricSpace/GromovHausdorff.lean b/Mathlib/Topology/MetricSpace/GromovHausdorff.lean index 8b58e2adf2ada..7082a492d3c8f 100644 --- a/Mathlib/Topology/MetricSpace/GromovHausdorff.lean +++ b/Mathlib/Topology/MetricSpace/GromovHausdorff.lean @@ -789,7 +789,7 @@ theorem totallyBounded {t : Set GHSpace} {C : ℝ} {u : ℕ → ℝ} {K : ℕ exact fun hp' => (hp hp').elim · rcases hcov _ (Set.not_not_mem.1 hp) n with ⟨s, ⟨scard, scover⟩⟩ rcases Cardinal.lt_aleph0.1 (lt_of_le_of_lt scard (Cardinal.nat_lt_aleph0 _)) with ⟨N, hN⟩ - rw [hN, Cardinal.natCast_le] at scard + rw [hN, Nat.cast_le] at scard have : #s = #(Fin N) := by rw [hN, Cardinal.mk_fin] cases' Quotient.exact this with E use s, N, scard, E From 6e6c4bebe658cfab06f6ab9ced117685a2e81ed2 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Wed, 23 Oct 2024 06:16:46 +0000 Subject: [PATCH 318/505] =?UTF-8?q?chore(SetTheory/Cardinal/Basic):=20`mem?= =?UTF-8?q?=5Frange=5Fof=5Fle=5Flift`=20=E2=86=92=20`mem=5Frange=5Flift=5F?= =?UTF-8?q?of=5Fle`=20(#17970)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This was a typo I made in #16958 only a week ago or so, so hopefully this doesn't warrant deprecations. --- Mathlib/SetTheory/Cardinal/Basic.lean | 8 ++++---- Mathlib/SetTheory/Cardinal/Cofinality.lean | 2 +- 2 files changed, 5 insertions(+), 5 deletions(-) diff --git a/Mathlib/SetTheory/Cardinal/Basic.lean b/Mathlib/SetTheory/Cardinal/Basic.lean index 10ebaa481ad46..cfd91460ac93e 100644 --- a/Mathlib/SetTheory/Cardinal/Basic.lean +++ b/Mathlib/SetTheory/Cardinal/Basic.lean @@ -335,14 +335,14 @@ def liftInitialSeg : Cardinal.{u} ≤i Cardinal.{max u v} := by rintro ⟨a, ⟨b, rfl⟩⟩ exact ⟨b, rfl⟩ -theorem mem_range_of_le_lift {a : Cardinal.{u}} {b : Cardinal.{max u v}} : +theorem mem_range_lift_of_le {a : Cardinal.{u}} {b : Cardinal.{max u v}} : b ≤ lift.{v, u} a → b ∈ Set.range lift.{v, u} := liftInitialSeg.mem_range_of_le -@[deprecated mem_range_of_le_lift (since := "2024-10-07")] +@[deprecated mem_range_lift_of_le (since := "2024-10-07")] theorem lift_down {a : Cardinal.{u}} {b : Cardinal.{max u v}} : b ≤ lift.{v, u} a → ∃ a', lift.{v, u} a' = b := - mem_range_of_le_lift + mem_range_lift_of_le /-- `Cardinal.lift` as an `OrderEmbedding`. -/ @[deprecated Cardinal.liftInitialSeg (since := "2024-10-07")] @@ -1037,7 +1037,7 @@ theorem lift_sSup {s : Set Cardinal} (hs : BddAbove s) : apply ((le_csSup_iff' (bddAbove_image.{_,u} _ hs)).2 fun c hc => _).antisymm (csSup_le' _) · intro c hc by_contra h - obtain ⟨d, rfl⟩ := Cardinal.mem_range_of_le_lift (not_le.1 h).le + obtain ⟨d, rfl⟩ := Cardinal.mem_range_lift_of_le (not_le.1 h).le simp_rw [lift_le] at h hc rw [csSup_le_iff' hs] at h exact h fun a ha => lift_le.1 <| hc (mem_image_of_mem _ ha) diff --git a/Mathlib/SetTheory/Cardinal/Cofinality.lean b/Mathlib/SetTheory/Cardinal/Cofinality.lean index e0ff450fad422..08369b5257281 100644 --- a/Mathlib/SetTheory/Cardinal/Cofinality.lean +++ b/Mathlib/SetTheory/Cardinal/Cofinality.lean @@ -700,7 +700,7 @@ theorem cof_univ : cof univ.{u, v} = Cardinal.univ.{u, v} := rcases @cof_eq Ordinal.{u} (· < ·) _ with ⟨S, H, Se⟩ rw [univ, ← lift_cof, ← Cardinal.lift_lift.{u+1, v, u}, Cardinal.lift_lt, ← Se] refine lt_of_not_ge fun h => ?_ - cases' Cardinal.mem_range_of_le_lift h with a e + cases' Cardinal.mem_range_lift_of_le h with a e refine Quotient.inductionOn a (fun α e => ?_) e cases' Quotient.exact e with f have f := Equiv.ulift.symm.trans f From 8164cd2e7afacdfa0016ae8d26b8f425232b49f7 Mon Sep 17 00:00:00 2001 From: Jz Pan Date: Wed, 23 Oct 2024 06:50:35 +0000 Subject: [PATCH 319/505] doc(CategoryTheory/Limits/Sifted): remove non-ASCII characters in bibliography key (#18101) Currently doc-gen4 does not allow non-ASCII characters in bibkey. --- Mathlib/CategoryTheory/Limits/Sifted.lean | 2 +- docs/references.bib | 22 +++++++++++----------- 2 files changed, 12 insertions(+), 12 deletions(-) diff --git a/Mathlib/CategoryTheory/Limits/Sifted.lean b/Mathlib/CategoryTheory/Limits/Sifted.lean index 0949b037550c1..8f5721d4f8dae 100644 --- a/Mathlib/CategoryTheory/Limits/Sifted.lean +++ b/Mathlib/CategoryTheory/Limits/Sifted.lean @@ -17,7 +17,7 @@ preserves finite products. ## References - [nLab, *Sifted category*](https://ncatlab.org/nlab/show/sifted+category) -- [*Algebraic Theories*, Chapter 2.][Adámek_Rosický_Vitale_2010] +- [*Algebraic Theories*, Chapter 2.][Adamek_Rosicky_Vitale_2010] -/ universe w v v₁ u u₁ diff --git a/docs/references.bib b/docs/references.bib index e94cd696fd232..f7bd7c609f40e 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -21,6 +21,17 @@ @Book{ abramsky_gabbay_maibaum_1994 zbl = {0829.68111} } +@Book{ Adamek_Rosicky_Vitale_2010, + place = {Cambridge}, + series = {Cambridge Tracts in Mathematics}, + title = {Algebraic Theories: A Categorical Introduction to General + Algebra}, + publisher = {Cambridge University Press}, + author = {Adámek, J. and Rosický, J. and Vitale, E. M.}, + year = {2010}, + collection = {Cambridge Tracts in Mathematics} +} + @InProceedings{ adhesive2004, author = {S. Lack and P. Soboci{\'n}ski}, title = {Adhesive categories}, @@ -34,17 +45,6 @@ @InProceedings{ adhesive2004 url = {https://www.ioc.ee/~pawel/papers/adhesive.pdf} } -@Book{ Adámek_Rosický_Vitale_2010, - place = {Cambridge}, - series = {Cambridge Tracts in Mathematics}, - title = {Algebraic Theories: A Categorical Introduction to General - Algebra}, - publisher = {Cambridge University Press}, - author = {Adámek, J. and Rosický, J. and Vitale, E. M.}, - year = {2010}, - collection = {Cambridge Tracts in Mathematics} -} - @Article{ ahrens2017, author = {Benedikt Ahrens and Peter LeFanu Lumsdaine}, year = {2019}, From 11c09b701eb2f5359b879b70b9caf521c24eced9 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Wed, 23 Oct 2024 07:27:44 +0000 Subject: [PATCH 320/505] feat: more tail recursion when finding Mersenne primes (#16168) Previously the failure at `(mersenne 9689).Prime` was a stack overflow in a `norm_num` helper function. Making the relevant function tail-recursive solves that. We now fail with `(kernel) deep recursion detected`, at an apparently system dependent step: * in CI, we still fail at `mersenne 9689` (just with a different error) * on my machine we now succeed at `9689` and `9941`, then fail at `11213`. --- Archive/Examples/MersennePrimes.lean | 20 +++++++-- Mathlib/NumberTheory/LucasLehmer.lean | 62 +++++++++++++++++++++------ 2 files changed, 67 insertions(+), 15 deletions(-) diff --git a/Archive/Examples/MersennePrimes.lean b/Archive/Examples/MersennePrimes.lean index a19f23f8431d4..ed395b0b58e41 100644 --- a/Archive/Examples/MersennePrimes.lean +++ b/Archive/Examples/MersennePrimes.lean @@ -80,8 +80,22 @@ example : (mersenne 4253).Prime := example : (mersenne 4423).Prime := lucas_lehmer_sufficiency _ (by norm_num) (by norm_num) --- First failure ("deep recursion detected") /- -example : (mersenne 9689).Prime := - lucas_lehmer_sufficiency _ (by norm_num) (by norm_num) +`mersenne 9689` seems to be system dependent: +locally it works fine, but in CI it fails with `(kernel) deep recursion detected` +-/ +-- example : (mersenne 9689).Prime := +-- lucas_lehmer_sufficiency _ (by norm_num) (by norm_num) + +/- +`mersenne 9941` seems to be system dependent: +locally it works fine, but in CI it fails with `(kernel) deep recursion detected` +-/ +-- example : (mersenne 9941).Prime := +-- lucas_lehmer_sufficiency _ (by norm_num) (by norm_num) + +/- +`mersenne 11213` fails with `(kernel) deep recursion detected` locally as well. -/ +-- example : (mersenne 11213).Prime := +-- lucas_lehmer_sufficiency _ (by norm_num) (by norm_num) diff --git a/Mathlib/NumberTheory/LucasLehmer.lean b/Mathlib/NumberTheory/LucasLehmer.lean index 91ce335ec0865..7990d0b3b0785 100644 --- a/Mathlib/NumberTheory/LucasLehmer.lean +++ b/Mathlib/NumberTheory/LucasLehmer.lean @@ -506,21 +506,21 @@ open Qq Lean Elab.Tactic Mathlib.Meta.NormNum /-- Version of `sMod` that is `ℕ`-valued. One should have `q = 2 ^ p - 1`. This can be reduced by the kernel. -/ -def sMod' (q : ℕ) : ℕ → ℕ +def sModNat (q : ℕ) : ℕ → ℕ | 0 => 4 % q - | i + 1 => (sMod' q i ^ 2 + (q - 2)) % q + | i + 1 => (sModNat q i ^ 2 + (q - 2)) % q -theorem sMod'_eq_sMod (p k : ℕ) (hp : 2 ≤ p) : (sMod' (2 ^ p - 1) k : ℤ) = sMod p k := by +theorem sModNat_eq_sMod (p k : ℕ) (hp : 2 ≤ p) : (sModNat (2 ^ p - 1) k : ℤ) = sMod p k := by have h1 := calc 4 = 2 ^ 2 := by norm_num _ ≤ 2 ^ p := Nat.pow_le_pow_of_le_right (by norm_num) hp have h2 : 1 ≤ 2 ^ p := by omega induction k with | zero => - rw [sMod', sMod, Int.ofNat_emod] + rw [sModNat, sMod, Int.ofNat_emod] simp [h2] | succ k ih => - rw [sMod', sMod, ← ih] + rw [sModNat, sMod, ← ih] have h3 : 2 ≤ 2 ^ p - 1 := by zify [h2] calc @@ -530,17 +530,55 @@ theorem sMod'_eq_sMod (p k : ℕ) (hp : 2 ≤ p) : (sMod' (2 ^ p - 1) k : ℤ) = rw [← add_sub_assoc, sub_eq_add_neg, add_assoc, add_comm _ (-2), ← add_assoc, Int.add_emod_self, ← sub_eq_add_neg] -lemma testTrueHelper (p : ℕ) (hp : Nat.blt 1 p = true) (h : sMod' (2 ^ p - 1) (p - 2) = 0) : +/-- Tail-recursive version of `sModNat`. -/ +def sModNatTR (q : ℕ) (k : Nat) : ℕ := + go k (4 % q) +where + /-- Helper function for `sMod''`. -/ + go : ℕ → ℕ → ℕ + | 0, acc => acc + | n + 1, acc => go n ((acc ^ 2 + (q - 2)) % q) + +/-- +Generalization of `sModNat` with arbitrary base case, +useful for proving `sModNatTR` and `sModNat` agree. +-/ +def sModNat_aux (b : ℕ) (q : ℕ) : ℕ → ℕ + | 0 => b + | i + 1 => (sModNat_aux b q i ^ 2 + (q - 2)) % q + +theorem sModNat_aux_eq (q k : ℕ) : sModNat_aux (4 % q) q k = sModNat q k := by + induction k with + | zero => rfl + | succ k ih => rw [sModNat_aux, ih, sModNat, ← ih] + +theorem sModNatTR_eq_sModNat (q : ℕ) (i : ℕ) : sModNatTR q i = sModNat q i := by + rw [sModNatTR, helper, sModNat_aux_eq] +where + helper b q k : sModNatTR.go q k b = sModNat_aux b q k := by + induction k generalizing b with + | zero => rfl + | succ k ih => + rw [sModNatTR.go, ih, sModNat_aux] + clear ih + induction k with + | zero => rfl + | succ k ih => + rw [sModNat_aux, ih, sModNat_aux] + +lemma testTrueHelper (p : ℕ) (hp : Nat.blt 1 p = true) (h : sModNatTR (2 ^ p - 1) (p - 2) = 0) : LucasLehmerTest p := by rw [Nat.blt_eq] at hp - rw [LucasLehmerTest, LucasLehmer.residue_eq_zero_iff_sMod_eq_zero p hp, ← sMod'_eq_sMod p _ hp, h] + rw [LucasLehmerTest, LucasLehmer.residue_eq_zero_iff_sMod_eq_zero p hp, ← sModNat_eq_sMod p _ hp, + ← sModNatTR_eq_sModNat, h] rfl lemma testFalseHelper (p : ℕ) (hp : Nat.blt 1 p = true) - (h : Nat.ble 1 (sMod' (2 ^ p - 1) (p - 2))) : ¬ LucasLehmerTest p := by + (h : Nat.ble 1 (sModNatTR (2 ^ p - 1) (p - 2))) : ¬ LucasLehmerTest p := by rw [Nat.blt_eq] at hp rw [Nat.ble_eq, Nat.succ_le, Nat.pos_iff_ne_zero] at h - rw [LucasLehmerTest, LucasLehmer.residue_eq_zero_iff_sMod_eq_zero p hp, ← sMod'_eq_sMod p _ hp] + rw [LucasLehmerTest, LucasLehmer.residue_eq_zero_iff_sMod_eq_zero p hp, ← sModNat_eq_sMod p _ hp, + ← sModNatTR_eq_sModNat] simpa using h theorem isNat_lucasLehmerTest : {p np : ℕ} → @@ -561,13 +599,13 @@ def evalLucasLehmerTest : NormNumExt where eval {_ _} e := do unless 1 < np do failure haveI' h1ltp : Nat.blt 1 $ep =Q true := ⟨⟩ - if sMod' (2 ^ np - 1) (np - 2) = 0 then - haveI' hs : sMod' (2 ^ $ep - 1) ($ep - 2) =Q 0 := ⟨⟩ + if sModNatTR (2 ^ np - 1) (np - 2) = 0 then + haveI' hs : sModNatTR (2 ^ $ep - 1) ($ep - 2) =Q 0 := ⟨⟩ have pf : Q(LucasLehmerTest $ep) := q(testTrueHelper $ep $h1ltp $hs) have pf' : Q(LucasLehmerTest $p) := q(isNat_lucasLehmerTest $hp $pf) return .isTrue pf' else - haveI' hs : Nat.ble 1 (sMod' (2 ^ $ep - 1) ($ep - 2)) =Q true := ⟨⟩ + haveI' hs : Nat.ble 1 (sModNatTR (2 ^ $ep - 1) ($ep - 2)) =Q true := ⟨⟩ have pf : Q(¬ LucasLehmerTest $ep) := q(testFalseHelper $ep $h1ltp $hs) have pf' : Q(¬ LucasLehmerTest $p) := q(isNat_not_lucasLehmerTest $hp $pf) return .isFalse pf' From 80a505f6a533d327c3f8bffadf511660a4be953d Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Wed, 23 Oct 2024 07:27:45 +0000 Subject: [PATCH 321/505] feat: generalize Algebra.TensorProduct.rightAlgebra (#18089) Co-authored-by: Kevin Buzzard --- Mathlib/RingTheory/TensorProduct/Basic.lean | 14 +++++++++++--- 1 file changed, 11 insertions(+), 3 deletions(-) diff --git a/Mathlib/RingTheory/TensorProduct/Basic.lean b/Mathlib/RingTheory/TensorProduct/Basic.lean index cbcde91fc95fd..dbb11acf62992 100644 --- a/Mathlib/RingTheory/TensorProduct/Basic.lean +++ b/Mathlib/RingTheory/TensorProduct/Basic.lean @@ -576,12 +576,22 @@ instance instCommRing : CommRing (A ⊗[R] B) := { toRing := inferInstance mul_comm := mul_comm } +end CommRing + section RightAlgebra +variable [CommSemiring R] +variable [Semiring A] [Algebra R A] +variable [CommSemiring B] [Algebra R B] + /-- `S ⊗[R] T` has a `T`-algebra structure. This is not a global instance or else the action of `S` on `S ⊗[R] S` would be ambiguous. -/ abbrev rightAlgebra : Algebra B (A ⊗[R] B) := - (Algebra.TensorProduct.includeRight.toRingHom : B →+* A ⊗[R] B).toAlgebra + includeRight.toRingHom.toAlgebra' fun b x => by + suffices LinearMap.mulLeft R (includeRight b) = LinearMap.mulRight R (includeRight b) from + congr($this x) + ext xa xb + simp [mul_comm] attribute [local instance] TensorProduct.rightAlgebra @@ -590,8 +600,6 @@ instance right_isScalarTower : IsScalarTower R B (A ⊗[R] B) := end RightAlgebra -end CommRing - /-- Verify that typeclass search finds the ring structure on `A ⊗[ℤ] B` when `A` and `B` are merely rings, by treating both as `ℤ`-algebras. -/ From 029a1d4c328dac22fa2d9cd510094d5b0fd5b636 Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Wed, 23 Oct 2024 07:52:49 +0000 Subject: [PATCH 322/505] feat: construct homomorphism from `MulSemiringAction` (#17849) This PR adds some missing API for constructing a homomorphism from a `MulSemiringAction`s. We already have `MulSemiringAction.toRingAut`. --- Mathlib/Algebra/Algebra/Equiv.lean | 10 ++++++++++ 1 file changed, 10 insertions(+) diff --git a/Mathlib/Algebra/Algebra/Equiv.lean b/Mathlib/Algebra/Algebra/Equiv.lean index 89b6ce8e80381..d1c80042e9fe6 100644 --- a/Mathlib/Algebra/Algebra/Equiv.lean +++ b/Mathlib/Algebra/Algebra/Equiv.lean @@ -801,6 +801,16 @@ theorem toAlgEquiv_injective [FaithfulSMul G A] : Function.Injective (MulSemiringAction.toAlgEquiv R A : G → A ≃ₐ[R] A) := fun _ _ h => eq_of_smul_eq_smul fun r => AlgEquiv.ext_iff.1 h r +/-- Each element of the group defines an algebra equivalence. + +This is a stronger version of `MulSemiringAction.toRingAut` and +`DistribMulAction.toModuleEnd`. -/ +@[simps] +def toAlgAut : G →* A ≃ₐ[R] A where + toFun := toAlgEquiv R A + map_one' := AlgEquiv.ext <| one_smul _ + map_mul' g h := AlgEquiv.ext <| mul_smul g h + end end MulSemiringAction From 2983bc58188940bb287d5e76b9150adc172ca2cd Mon Sep 17 00:00:00 2001 From: "Tristan F.-R." Date: Wed, 23 Oct 2024 08:52:47 +0000 Subject: [PATCH 323/505] chore(Wiedijk/InverseTriangleSum): golf (#18100) Since `InverseTriangleSum`'s original writing, `field_simp` has become stronger. Co-authored-by: Tristan F. --- Archive/Wiedijk100Theorems/InverseTriangleSum.lean | 2 -- 1 file changed, 2 deletions(-) diff --git a/Archive/Wiedijk100Theorems/InverseTriangleSum.lean b/Archive/Wiedijk100Theorems/InverseTriangleSum.lean index 6a32da4b6aa4c..e1ec0f91f2648 100644 --- a/Archive/Wiedijk100Theorems/InverseTriangleSum.lean +++ b/Archive/Wiedijk100Theorems/InverseTriangleSum.lean @@ -30,7 +30,5 @@ theorem Theorems100.inverse_triangle_sum : refine sum_range_induction _ _ (if_pos rfl) ?_ rintro (_ | n) · rw [if_neg, if_pos] <;> norm_num - simp only [Nat.succ_ne_zero, ↓reduceIte, Nat.cast_succ] - have A : (n + 1 + 1 : ℚ) ≠ 0 := by norm_cast field_simp ring From 52e1bb93269e67b879815db62edb5e00cc626108 Mon Sep 17 00:00:00 2001 From: Dagur Asgeirsson Date: Wed, 23 Oct 2024 09:21:54 +0000 Subject: [PATCH 324/505] chore(CategoryTheory/Sites): remove a resolved TODO (#18104) --- Mathlib/CategoryTheory/Sites/ConstantSheaf.lean | 5 ----- 1 file changed, 5 deletions(-) diff --git a/Mathlib/CategoryTheory/Sites/ConstantSheaf.lean b/Mathlib/CategoryTheory/Sites/ConstantSheaf.lean index 46db78d658c1f..a300122560b4c 100644 --- a/Mathlib/CategoryTheory/Sites/ConstantSheaf.lean +++ b/Mathlib/CategoryTheory/Sites/ConstantSheaf.lean @@ -27,11 +27,6 @@ under equivalence of sheaf categories. * `Sheaf.isConstant_iff_forget` : Given a "forgetful" functor `U : D ⥤ B` a sheaf `F : Sheaf J D` is constant if and only if the sheaf given by postcomposition with `U` is constant. - -## Future work - -* (Dagur) Use `Sheaf.isConstant_iff_forget` to prove that a condensed module is discrete if and -only if its underlying condensed set is discrete. -/ namespace CategoryTheory From 5a2f1ac951b5455e1863c4efa7c6f5a01981218c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Wed, 23 Oct 2024 09:32:35 +0000 Subject: [PATCH 325/505] =?UTF-8?q?refactor(SetTheory/Cardinal/Aleph):=20`?= =?UTF-8?q?aleph'`=20=E2=86=92=20`preAleph`=20(#18099)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This helps us move away from primes in theorems and definitions. Suggested by Floris van Doorn. --- Mathlib/SetTheory/Cardinal/Aleph.lean | 353 ++++++++++++--------- Mathlib/SetTheory/Cardinal/Cofinality.lean | 11 + 2 files changed, 217 insertions(+), 147 deletions(-) diff --git a/Mathlib/SetTheory/Cardinal/Aleph.lean b/Mathlib/SetTheory/Cardinal/Aleph.lean index c1473e55bb543..a7267f649efde 100644 --- a/Mathlib/SetTheory/Cardinal/Aleph.lean +++ b/Mathlib/SetTheory/Cardinal/Aleph.lean @@ -10,8 +10,8 @@ import Mathlib.SetTheory.Ordinal.Arithmetic /-! # Aleph and beth functions -* The function `Cardinal.aleph'` gives the cardinals listed by their ordinal index. - `aleph' n = n`, `aleph' ω = ℵ₀`, `aleph' (ω + 1) = succ ℵ₀`, etc. +* The function `Cardinal.preAleph` gives the cardinals listed by their ordinal index. + `preAleph n = n`, `preAleph ω = ℵ₀`, `preAleph (ω + 1) = succ ℵ₀`, etc. It is an order isomorphism between ordinals and cardinals. * The function `Cardinal.aleph` gives the infinite cardinals listed by their ordinal index. `aleph 0 = ℵ₀`, `aleph 1 = succ ℵ₀` is the first @@ -100,11 +100,11 @@ end Ordinal namespace Cardinal -/-- The `aleph'` function gives the cardinals listed by their ordinal index. `aleph' n = n`, -`aleph' ω = ℵ₀`, `aleph' (ω + 1) = succ ℵ₀`, etc. +/-- The "pre-aleph" function gives the cardinals listed by their ordinal index. `preAleph n = n`, +`preAleph ω = ℵ₀`, `preAleph (ω + 1) = succ ℵ₀`, etc. For the more common aleph function skipping over finite cardinals, see `Cardinal.aleph`. -/ -def aleph' : Ordinal.{u} ≃o Cardinal.{u} := by +def preAleph : Ordinal.{u} ≃o Cardinal.{u} := by let f := RelEmbedding.collapse Cardinal.ord.orderEmbedding.ltEmbedding.{u} refine (OrderIso.ofRelIsoLT <| RelIso.ofSurjective f ?_).symm apply f.eq_or_principal.resolve_right @@ -120,159 +120,74 @@ def aleph' : Ordinal.{u} ≃o Cardinal.{u} := by (Cardinal.bddAbove_range.{u, u} fun a : α => invFun f (Ordinal.typein r a)) (Ordinal.enum r ⟨_, h (succ s)⟩) -/-- The `aleph'` index function, which gives the ordinal index of a cardinal. - (The `aleph'` part is because unlike `aleph` this counts also the - finite stages. So `alephIdx n = n`, `alephIdx ω = ω`, - `alephIdx ℵ₁ = ω + 1` and so on.) - In this definition, we register additionally that this function is an initial segment, - i.e., it is order preserving and its range is an initial segment of the ordinals. - For the basic function version, see `alephIdx`. - For an upgraded version stating that the range is everything, see `AlephIdx.rel_iso`. -/ -@[deprecated aleph' (since := "2024-08-28")] -def alephIdx.initialSeg : @InitialSeg Cardinal Ordinal (· < ·) (· < ·) := - @RelEmbedding.collapse Cardinal Ordinal (· < ·) (· < ·) _ Cardinal.ord.orderEmbedding.ltEmbedding - -/-- The `aleph'` index function, which gives the ordinal index of a cardinal. - (The `aleph'` part is because unlike `aleph` this counts also the - finite stages. So `alephIdx n = n`, `alephIdx ℵ₀ = ω`, - `alephIdx ℵ₁ = ω + 1` and so on.) - In this version, we register additionally that this function is an order isomorphism - between cardinals and ordinals. - For the basic function version, see `alephIdx`. -/ -@[deprecated aleph' (since := "2024-08-28")] -def alephIdx.relIso : @RelIso Cardinal.{u} Ordinal.{u} (· < ·) (· < ·) := - aleph'.symm.toRelIsoLT - -/-- The `aleph'` index function, which gives the ordinal index of a cardinal. - (The `aleph'` part is because unlike `aleph` this counts also the - finite stages. So `alephIdx n = n`, `alephIdx ω = ω`, - `alephIdx ℵ₁ = ω + 1` and so on.) - For an upgraded version stating that the range is everything, see `AlephIdx.rel_iso`. -/ -@[deprecated aleph' (since := "2024-08-28")] -def alephIdx : Cardinal → Ordinal := - aleph'.symm - -set_option linter.deprecated false in -@[deprecated (since := "2024-08-28")] -theorem alephIdx.initialSeg_coe : (alephIdx.initialSeg : Cardinal → Ordinal) = alephIdx := - rfl - -set_option linter.deprecated false in -@[deprecated (since := "2024-08-28")] -theorem alephIdx_lt {a b} : alephIdx a < alephIdx b ↔ a < b := - alephIdx.initialSeg.toRelEmbedding.map_rel_iff - -set_option linter.deprecated false in -@[deprecated (since := "2024-08-28")] -theorem alephIdx_le {a b} : alephIdx a ≤ alephIdx b ↔ a ≤ b := by - rw [← not_lt, ← not_lt, alephIdx_lt] - -set_option linter.deprecated false in -@[deprecated (since := "2024-08-28")] -theorem alephIdx.init {a b} : b < alephIdx a → ∃ c, alephIdx c = b := - alephIdx.initialSeg.init - -set_option linter.deprecated false in -@[deprecated (since := "2024-08-28")] -theorem alephIdx.relIso_coe : (alephIdx.relIso : Cardinal → Ordinal) = alephIdx := - rfl - @[simp] theorem type_cardinal : @type Cardinal (· < ·) _ = Ordinal.univ.{u, u + 1} := by rw [Ordinal.univ_id] - exact Quotient.sound ⟨aleph'.symm.toRelIsoLT⟩ + exact Quotient.sound ⟨preAleph.symm.toRelIsoLT⟩ @[simp] theorem mk_cardinal : #Cardinal = univ.{u, u + 1} := by simpa only [card_type, card_univ] using congr_arg card type_cardinal -/-- The `aleph'` function gives the cardinals listed by their ordinal - index, and is the inverse of `aleph_idx`. - `aleph' n = n`, `aleph' ω = ω`, `aleph' (ω + 1) = succ ℵ₀`, etc. - In this version, we register additionally that this function is an order isomorphism - between ordinals and cardinals. - For the basic function version, see `aleph'`. -/ -@[deprecated aleph' (since := "2024-08-28")] -def Aleph'.relIso := - aleph' +theorem preAleph_lt_preAleph {o₁ o₂ : Ordinal} : preAleph o₁ < preAleph o₂ ↔ o₁ < o₂ := + preAleph.lt_iff_lt -set_option linter.deprecated false in -@[deprecated (since := "2024-08-28")] -theorem aleph'.relIso_coe : (Aleph'.relIso : Ordinal → Cardinal) = aleph' := - rfl +theorem preAleph_le_preAleph {o₁ o₂ : Ordinal} : preAleph o₁ ≤ preAleph o₂ ↔ o₁ ≤ o₂ := + preAleph.le_iff_le -theorem aleph'_lt {o₁ o₂ : Ordinal} : aleph' o₁ < aleph' o₂ ↔ o₁ < o₂ := - aleph'.lt_iff_lt - -theorem aleph'_le {o₁ o₂ : Ordinal} : aleph' o₁ ≤ aleph' o₂ ↔ o₁ ≤ o₂ := - aleph'.le_iff_le - -theorem aleph'_max (o₁ o₂ : Ordinal) : aleph' (max o₁ o₂) = max (aleph' o₁) (aleph' o₂) := - aleph'.monotone.map_max - -set_option linter.deprecated false in -@[deprecated (since := "2024-08-28")] -theorem aleph'_alephIdx (c : Cardinal) : aleph' c.alephIdx = c := - Cardinal.alephIdx.relIso.toEquiv.symm_apply_apply c - -set_option linter.deprecated false in -@[deprecated (since := "2024-08-28")] -theorem alephIdx_aleph' (o : Ordinal) : (aleph' o).alephIdx = o := - Cardinal.alephIdx.relIso.toEquiv.apply_symm_apply o +theorem preAleph_max (o₁ o₂ : Ordinal) : preAleph (max o₁ o₂) = max (preAleph o₁) (preAleph o₂) := + preAleph.monotone.map_max @[simp] -theorem aleph'_zero : aleph' 0 = 0 := - aleph'.map_bot +theorem preAleph_zero : preAleph 0 = 0 := + preAleph.map_bot @[simp] -theorem aleph'_succ (o : Ordinal) : aleph' (succ o) = succ (aleph' o) := - aleph'.map_succ o +theorem preAleph_succ (o : Ordinal) : preAleph (succ o) = succ (preAleph o) := + preAleph.map_succ o @[simp] -theorem aleph'_nat : ∀ n : ℕ, aleph' n = n - | 0 => aleph'_zero - | n + 1 => show aleph' (succ n) = n.succ by rw [aleph'_succ, aleph'_nat n, nat_succ] +theorem preAleph_nat : ∀ n : ℕ, preAleph n = n + | 0 => preAleph_zero + | n + 1 => show preAleph (succ n) = n.succ by rw [preAleph_succ, preAleph_nat n, nat_succ] -set_option linter.docPrime false in -@[simp] -theorem lift_aleph' (o : Ordinal.{u}) : lift.{v} (aleph' o) = aleph' (Ordinal.lift.{v} o) := - ((InitialSeg.ofIso aleph'.toRelIsoLT).trans liftInitialSeg).eq - (Ordinal.liftInitialSeg.trans (InitialSeg.ofIso aleph'.toRelIsoLT)) o +theorem preAleph_pos {o : Ordinal} : 0 < preAleph o ↔ 0 < o := by + rw [← preAleph_zero, preAleph_lt_preAleph] -theorem aleph'_le_of_limit {o : Ordinal} (l : o.IsLimit) {c} : - aleph' o ≤ c ↔ ∀ o' < o, aleph' o' ≤ c := - ⟨fun h o' h' => (aleph'_le.2 <| h'.le).trans h, fun h => by - rw [← aleph'.apply_symm_apply c, aleph'_le, limit_le l] +@[simp] +theorem lift_preAleph (o : Ordinal.{u}) : lift.{v} (preAleph o) = preAleph (Ordinal.lift.{v} o) := + ((InitialSeg.ofIso preAleph.toRelIsoLT).trans liftInitialSeg).eq + (Ordinal.liftInitialSeg.trans (InitialSeg.ofIso preAleph.toRelIsoLT)) o + +theorem preAleph_le_of_isLimit {o : Ordinal} (l : o.IsLimit) {c} : + preAleph o ≤ c ↔ ∀ o' < o, preAleph o' ≤ c := + ⟨fun h o' h' => (preAleph_le_preAleph.2 <| h'.le).trans h, fun h => by + rw [← preAleph.apply_symm_apply c, preAleph_le_preAleph, limit_le l] intro x h' - rw [← aleph'_le, aleph'.apply_symm_apply] + rw [← preAleph_le_preAleph, preAleph.apply_symm_apply] exact h _ h'⟩ -theorem aleph'_limit {o : Ordinal} (ho : o.IsLimit) : aleph' o = ⨆ a : Iio o, aleph' a := by - refine le_antisymm ?_ (ciSup_le' fun i => aleph'_le.2 (le_of_lt i.2)) - rw [aleph'_le_of_limit ho] +theorem preAleph_limit {o : Ordinal} (ho : o.IsLimit) : preAleph o = ⨆ a : Iio o, preAleph a := by + refine le_antisymm ?_ (ciSup_le' fun i => preAleph_le_preAleph.2 i.2.le) + rw [preAleph_le_of_isLimit ho] exact fun a ha => le_ciSup (bddAbove_of_small _) (⟨a, ha⟩ : Iio o) @[simp] -theorem aleph'_omega0 : aleph' ω = ℵ₀ := +theorem preAleph_omega0 : preAleph ω = ℵ₀ := eq_of_forall_ge_iff fun c => by - simp only [aleph'_le_of_limit isLimit_omega0, lt_omega0, exists_imp, aleph0_le] - exact forall_swap.trans (forall_congr' fun n => by simp only [forall_eq, aleph'_nat]) - -@[deprecated (since := "2024-09-30")] -alias aleph'_omega := aleph'_omega0 + simp only [preAleph_le_of_isLimit isLimit_omega0, lt_omega0, exists_imp, aleph0_le] + exact forall_swap.trans (forall_congr' fun n => by simp only [forall_eq, preAleph_nat]) -set_option linter.deprecated false in -/-- `aleph'` and `aleph_idx` form an equivalence between `Ordinal` and `Cardinal` -/ -@[deprecated aleph' (since := "2024-08-28")] -def aleph'Equiv : Ordinal ≃ Cardinal := - ⟨aleph', alephIdx, alephIdx_aleph', aleph'_alephIdx⟩ +@[simp] +theorem aleph0_le_preAleph {o : Ordinal} : ℵ₀ ≤ preAleph o ↔ ω ≤ o := by + rw [← preAleph_omega0, preAleph_le_preAleph] /-- The `aleph` function gives the infinite cardinals listed by their ordinal index. `aleph 0 = ℵ₀`, `aleph 1 = succ ℵ₀` is the first uncountable cardinal, and so on. For a version including finite cardinals, see `Cardinal.aleph'`. -/ def aleph : Ordinal ↪o Cardinal := - (OrderEmbedding.addLeft ω).trans aleph'.toOrderEmbedding + (OrderEmbedding.addLeft ω).trans preAleph.toOrderEmbedding @[inherit_doc] scoped notation "ℵ_ " => aleph @@ -280,15 +195,21 @@ scoped notation "ℵ_ " => aleph /-- `ℵ₁` is the first uncountable ordinal. -/ scoped notation "ℵ₁" => ℵ_ 1 -theorem aleph_eq_aleph' (o : Ordinal) : ℵ_ o = aleph' (ω + o) := +theorem aleph_eq_preAleph (o : Ordinal) : ℵ_ o = preAleph (ω + o) := rfl -theorem aleph_lt {o₁ o₂ : Ordinal} : ℵ_ o₁ < ℵ_ o₂ ↔ o₁ < o₂ := +theorem aleph_lt_aleph {o₁ o₂ : Ordinal} : ℵ_ o₁ < ℵ_ o₂ ↔ o₁ < o₂ := aleph.lt_iff_lt -theorem aleph_le {o₁ o₂ : Ordinal} : ℵ_ o₁ ≤ ℵ_ o₂ ↔ o₁ ≤ o₂ := +@[deprecated aleph_lt_aleph (since := "2024-10-22")] +alias aleph_lt := aleph_lt_aleph + +theorem aleph_le_aleph {o₁ o₂ : Ordinal} : ℵ_ o₁ ≤ ℵ_ o₂ ↔ o₁ ≤ o₂ := aleph.le_iff_le +@[deprecated aleph_le_aleph (since := "2024-10-22")] +alias aleph_le := aleph_le_aleph + theorem aleph_max (o₁ o₂ : Ordinal) : ℵ_ (max o₁ o₂) = max (ℵ_ o₁) (ℵ_ o₂) := aleph.monotone.map_max @@ -298,35 +219,31 @@ theorem max_aleph_eq (o₁ o₂ : Ordinal) : max (ℵ_ o₁) (ℵ_ o₂) = ℵ_ @[simp] theorem aleph_succ (o : Ordinal) : ℵ_ (succ o) = succ (ℵ_ o) := by - rw [aleph_eq_aleph', add_succ, aleph'_succ, aleph_eq_aleph'] + rw [aleph_eq_preAleph, add_succ, preAleph_succ, aleph_eq_preAleph] @[simp] -theorem aleph_zero : ℵ_ 0 = ℵ₀ := by rw [aleph_eq_aleph', add_zero, aleph'_omega0] +theorem aleph_zero : ℵ_ 0 = ℵ₀ := by rw [aleph_eq_preAleph, add_zero, preAleph_omega0] @[simp] theorem lift_aleph (o : Ordinal.{u}) : lift.{v} (aleph o) = aleph (Ordinal.lift.{v} o) := by - simp [aleph_eq_aleph'] + simp [aleph_eq_preAleph] theorem aleph_limit {o : Ordinal} (ho : o.IsLimit) : ℵ_ o = ⨆ a : Iio o, ℵ_ a := by apply le_antisymm _ (ciSup_le' _) - · rw [aleph_eq_aleph', aleph'_limit (ho.add _)] + · rw [aleph_eq_preAleph, preAleph_limit (ho.add _)] refine ciSup_mono' (bddAbove_of_small _) ?_ rintro ⟨i, hi⟩ cases' lt_or_le i ω with h h · rcases lt_omega0.1 h with ⟨n, rfl⟩ use ⟨0, ho.pos⟩ simpa using (nat_lt_aleph0 n).le - · exact ⟨⟨_, (sub_lt_of_le h).2 hi⟩, aleph'_le.2 (le_add_sub _ _)⟩ - · exact fun i => aleph_le.2 (le_of_lt i.2) - -theorem aleph0_le_aleph' {o : Ordinal} : ℵ₀ ≤ aleph' o ↔ ω ≤ o := by rw [← aleph'_omega0, aleph'_le] + · exact ⟨⟨_, (sub_lt_of_le h).2 hi⟩, preAleph_le_preAleph.2 (le_add_sub _ _)⟩ + · exact fun i => aleph_le_aleph.2 i.2.le theorem aleph0_le_aleph (o : Ordinal) : ℵ₀ ≤ ℵ_ o := by - rw [aleph_eq_aleph', aleph0_le_aleph'] + rw [aleph_eq_preAleph, aleph0_le_preAleph] apply Ordinal.le_add_right -theorem aleph'_pos {o : Ordinal} (ho : 0 < o) : 0 < aleph' o := by rwa [← aleph'_zero, aleph'_lt] - theorem aleph_pos (o : Ordinal) : 0 < ℵ_ o := aleph0_pos.trans_le (aleph0_le_aleph o) @@ -350,17 +267,17 @@ instance (o : Ordinal) : NoMaxOrder (ℵ_ o).ord.toType := theorem exists_aleph {c : Cardinal} : ℵ₀ ≤ c ↔ ∃ o, c = ℵ_ o := ⟨fun h => - ⟨aleph'.symm c - ω, by - rw [aleph_eq_aleph', Ordinal.add_sub_cancel_of_le, aleph'.apply_symm_apply] - rwa [← aleph0_le_aleph', aleph'.apply_symm_apply]⟩, + ⟨preAleph.symm c - ω, by + rw [aleph_eq_preAleph, Ordinal.add_sub_cancel_of_le, preAleph.apply_symm_apply] + rwa [← aleph0_le_preAleph, preAleph.apply_symm_apply]⟩, fun ⟨o, e⟩ => e.symm ▸ aleph0_le_aleph _⟩ -theorem aleph'_isNormal : IsNormal (ord ∘ aleph') := - ⟨fun o => ord_lt_ord.2 <| aleph'_lt.2 <| lt_succ o, fun o l a => by - simp [ord_le, aleph'_le_of_limit l]⟩ +theorem preAleph_isNormal : IsNormal (ord ∘ preAleph) := + ⟨fun o => ord_lt_ord.2 <| preAleph_lt_preAleph.2 <| lt_succ o, fun o l a => by + simp [ord_le, preAleph_le_of_isLimit l]⟩ theorem aleph_isNormal : IsNormal (ord ∘ aleph) := - aleph'_isNormal.trans <| isNormal_add_right ω + preAleph_isNormal.trans <| isNormal_add_right ω theorem succ_aleph0 : succ ℵ₀ = ℵ₁ := by rw [← aleph_zero, ← aleph_succ, Ordinal.succ_zero] @@ -398,6 +315,148 @@ theorem lift_eq_aleph1 {c : Cardinal.{u}} : lift.{v} c = ℵ₁ ↔ c = ℵ₁ : section deprecated set_option linter.deprecated false +set_option linter.docPrime false + +@[deprecated preAleph (since := "2024-10-22")] +noncomputable alias aleph' := preAleph + +/-- The `aleph'` index function, which gives the ordinal index of a cardinal. + (The `aleph'` part is because unlike `aleph` this counts also the + finite stages. So `alephIdx n = n`, `alephIdx ω = ω`, + `alephIdx ℵ₁ = ω + 1` and so on.) + In this definition, we register additionally that this function is an initial segment, + i.e., it is order preserving and its range is an initial segment of the ordinals. + For the basic function version, see `alephIdx`. + For an upgraded version stating that the range is everything, see `AlephIdx.rel_iso`. -/ +@[deprecated preAleph (since := "2024-08-28")] +def alephIdx.initialSeg : @InitialSeg Cardinal Ordinal (· < ·) (· < ·) := + @RelEmbedding.collapse Cardinal Ordinal (· < ·) (· < ·) _ Cardinal.ord.orderEmbedding.ltEmbedding + +/-- The `aleph'` index function, which gives the ordinal index of a cardinal. + (The `aleph'` part is because unlike `aleph` this counts also the + finite stages. So `alephIdx n = n`, `alephIdx ℵ₀ = ω`, + `alephIdx ℵ₁ = ω + 1` and so on.) + In this version, we register additionally that this function is an order isomorphism + between cardinals and ordinals. + For the basic function version, see `alephIdx`. -/ +@[deprecated preAleph (since := "2024-08-28")] +def alephIdx.relIso : @RelIso Cardinal.{u} Ordinal.{u} (· < ·) (· < ·) := + aleph'.symm.toRelIsoLT + +/-- The `aleph'` index function, which gives the ordinal index of a cardinal. + (The `aleph'` part is because unlike `aleph` this counts also the + finite stages. So `alephIdx n = n`, `alephIdx ω = ω`, + `alephIdx ℵ₁ = ω + 1` and so on.) + For an upgraded version stating that the range is everything, see `AlephIdx.rel_iso`. -/ +@[deprecated aleph' (since := "2024-08-28")] +def alephIdx : Cardinal → Ordinal := + aleph'.symm + +@[deprecated (since := "2024-08-28")] +theorem alephIdx.initialSeg_coe : (alephIdx.initialSeg : Cardinal → Ordinal) = alephIdx := + rfl + +@[deprecated (since := "2024-08-28")] +theorem alephIdx_lt {a b} : alephIdx a < alephIdx b ↔ a < b := + alephIdx.initialSeg.toRelEmbedding.map_rel_iff + +@[deprecated (since := "2024-08-28")] +theorem alephIdx_le {a b} : alephIdx a ≤ alephIdx b ↔ a ≤ b := by + rw [← not_lt, ← not_lt, alephIdx_lt] + +@[deprecated (since := "2024-08-28")] +theorem alephIdx.init {a b} : b < alephIdx a → ∃ c, alephIdx c = b := + alephIdx.initialSeg.init + +@[deprecated (since := "2024-08-28")] +theorem alephIdx.relIso_coe : (alephIdx.relIso : Cardinal → Ordinal) = alephIdx := + rfl + +/-- The `aleph'` function gives the cardinals listed by their ordinal + index, and is the inverse of `aleph_idx`. + `aleph' n = n`, `aleph' ω = ω`, `aleph' (ω + 1) = succ ℵ₀`, etc. + In this version, we register additionally that this function is an order isomorphism + between ordinals and cardinals. + For the basic function version, see `aleph'`. -/ +@[deprecated aleph' (since := "2024-08-28")] +def Aleph'.relIso := + aleph' + +@[deprecated (since := "2024-08-28")] +theorem aleph'.relIso_coe : (Aleph'.relIso : Ordinal → Cardinal) = aleph' := + rfl + +@[deprecated preAleph_lt_preAleph (since := "2024-10-22")] +theorem aleph'_lt {o₁ o₂ : Ordinal} : aleph' o₁ < aleph' o₂ ↔ o₁ < o₂ := + aleph'.lt_iff_lt + +@[deprecated preAleph_le_preAleph (since := "2024-10-22")] +theorem aleph'_le {o₁ o₂ : Ordinal} : aleph' o₁ ≤ aleph' o₂ ↔ o₁ ≤ o₂ := + aleph'.le_iff_le + +@[deprecated preAleph_max (since := "2024-10-22")] +theorem aleph'_max (o₁ o₂ : Ordinal) : aleph' (max o₁ o₂) = max (aleph' o₁) (aleph' o₂) := + aleph'.monotone.map_max + +@[deprecated (since := "2024-08-28")] +theorem aleph'_alephIdx (c : Cardinal) : aleph' c.alephIdx = c := + Cardinal.alephIdx.relIso.toEquiv.symm_apply_apply c + +@[deprecated (since := "2024-08-28")] +theorem alephIdx_aleph' (o : Ordinal) : (aleph' o).alephIdx = o := + Cardinal.alephIdx.relIso.toEquiv.apply_symm_apply o + +@[deprecated preAleph_zero (since := "2024-10-22")] +theorem aleph'_zero : aleph' 0 = 0 := + aleph'.map_bot + +@[deprecated preAleph_succ (since := "2024-10-22")] +theorem aleph'_succ (o : Ordinal) : aleph' (succ o) = succ (aleph' o) := + aleph'.map_succ o + +@[deprecated preAleph_nat (since := "2024-10-22")] +theorem aleph'_nat : ∀ n : ℕ, aleph' n = n := + preAleph_nat + +@[deprecated lift_preAleph (since := "2024-10-22")] +theorem lift_aleph' (o : Ordinal.{u}) : lift.{v} (aleph' o) = aleph' (Ordinal.lift.{v} o) := + lift_preAleph o + +@[deprecated preAleph_le_of_isLimit (since := "2024-10-22")] +theorem aleph'_le_of_limit {o : Ordinal} (l : o.IsLimit) {c} : + aleph' o ≤ c ↔ ∀ o' < o, aleph' o' ≤ c := + preAleph_le_of_isLimit l + +@[deprecated preAleph_limit (since := "2024-10-22")] +theorem aleph'_limit {o : Ordinal} (ho : o.IsLimit) : aleph' o = ⨆ a : Iio o, aleph' a := + preAleph_limit ho + +@[deprecated preAleph_omega0 (since := "2024-10-22")] +theorem aleph'_omega0 : aleph' ω = ℵ₀ := + preAleph_omega0 + +@[deprecated (since := "2024-09-30")] +alias aleph'_omega := aleph'_omega0 + +/-- `aleph'` and `aleph_idx` form an equivalence between `Ordinal` and `Cardinal` -/ +@[deprecated aleph' (since := "2024-08-28")] +def aleph'Equiv : Ordinal ≃ Cardinal := + ⟨aleph', alephIdx, alephIdx_aleph', aleph'_alephIdx⟩ + +theorem aleph_eq_aleph' (o : Ordinal) : ℵ_ o = preAleph (ω + o) := + rfl + +@[deprecated aleph0_le_preAleph (since := "2024-10-22")] +theorem aleph0_le_aleph' {o : Ordinal} : ℵ₀ ≤ aleph' o ↔ ω ≤ o := by + rw [← aleph'_omega0, aleph'_le] + +@[deprecated preAleph_pos (since := "2024-10-22")] +theorem aleph'_pos {o : Ordinal} (ho : 0 < o) : 0 < aleph' o := by + rwa [← aleph'_zero, aleph'_lt] + +@[deprecated preAleph_isNormal (since := "2024-10-22")] +theorem aleph'_isNormal : IsNormal (ord ∘ aleph') := + preAleph_isNormal -- TODO: these lemmas should be stated in terms of the `ω` function and of an `IsInitial` predicate, -- neither of which currently exist. diff --git a/Mathlib/SetTheory/Cardinal/Cofinality.lean b/Mathlib/SetTheory/Cardinal/Cofinality.lean index 08369b5257281..9c31bb5780ec0 100644 --- a/Mathlib/SetTheory/Cardinal/Cofinality.lean +++ b/Mathlib/SetTheory/Cardinal/Cofinality.lean @@ -660,6 +660,11 @@ theorem aleph0_le_cof {o} : ℵ₀ ≤ cof o ↔ IsLimit o := by exact not_succ_isLimit _ l @[simp] +theorem preAleph_cof {o : Ordinal} (ho : o.IsLimit) : (preAleph o).ord.cof = o.cof := + preAleph_isNormal.cof_eq ho + +set_option linter.deprecated false in +@[deprecated preAleph_cof (since := "2024-10-22")] theorem aleph'_cof {o : Ordinal} (ho : o.IsLimit) : (aleph' o).ord.cof = o.cof := aleph'_isNormal.cof_eq ho @@ -924,6 +929,12 @@ theorem isRegular_aleph_one : IsRegular ℵ₁ := by rw [← succ_aleph0] exact isRegular_succ le_rfl +theorem isRegular_preAleph_succ {o : Ordinal} (h : ω ≤ o) : IsRegular (preAleph (succ o)) := by + rw [preAleph_succ] + exact isRegular_succ (aleph0_le_preAleph.2 h) + +set_option linter.deprecated false in +@[deprecated isRegular_preAleph_succ (since := "2024-10-22")] theorem isRegular_aleph'_succ {o : Ordinal} (h : ω ≤ o) : IsRegular (aleph' (succ o)) := by rw [aleph'_succ] exact isRegular_succ (aleph0_le_aleph'.2 h) From 9389838e744bd326e9b82a3ebd59df49c5b78af1 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Wed, 23 Oct 2024 10:42:15 +0000 Subject: [PATCH 326/505] chore(Algebra/Polynomial): `normalize_monic` as alias for `Monic.normalize_eq_self` (#18024) Discovered in #18006. As far as I can tell, neither lemma actually is used anywhere in Mathlib, so I arbitrarily picked the unnamespaced one to be a deprecated alias for the other. Please feel free to bikeshed over which one should win. --- Mathlib/Algebra/Polynomial/FieldDivision.lean | 5 +++-- 1 file changed, 3 insertions(+), 2 deletions(-) diff --git a/Mathlib/Algebra/Polynomial/FieldDivision.lean b/Mathlib/Algebra/Polynomial/FieldDivision.lean index 6a82a21dcee8f..f59c7998f4242 100644 --- a/Mathlib/Algebra/Polynomial/FieldDivision.lean +++ b/Mathlib/Algebra/Polynomial/FieldDivision.lean @@ -202,6 +202,9 @@ theorem Monic.normalize_eq_self {p : R[X]} (hp : p.Monic) : normalize p = p := b simp only [Polynomial.coe_normUnit, normalize_apply, hp.leadingCoeff, normUnit_one, Units.val_one, Polynomial.C.map_one, mul_one] +@[deprecated Polynomial.Monic.normalize_eq_self (since := "2024-10-21")] +alias normalize_monic := Monic.normalize_eq_self + theorem roots_normalize {p : R[X]} : (normalize p).roots = p.roots := by rw [normalize_apply, mul_comm, coe_normUnit, roots_C_mul _ (normUnit (leadingCoeff p)).ne_zero] @@ -524,8 +527,6 @@ theorem coe_normUnit_of_ne_zero [DecidableEq R] (hp : p ≠ 0) : have : p.leadingCoeff ≠ 0 := mt leadingCoeff_eq_zero.mp hp simp [CommGroupWithZero.coe_normUnit _ this] -theorem normalize_monic [DecidableEq R] (h : Monic p) : normalize p = p := by simp [h] - theorem map_dvd_map' [Field k] (f : R →+* k) {x y : R[X]} : x.map f ∣ y.map f ↔ x ∣ y := by by_cases H : x = 0 · rw [H, Polynomial.map_zero, zero_dvd_iff, zero_dvd_iff, map_eq_zero] From dcd264a67a3db96fcda1a6edd459549d08a2f30b Mon Sep 17 00:00:00 2001 From: Jiang Jiedong <107380768+jjdishere@users.noreply.github.com> Date: Wed, 23 Oct 2024 11:27:03 +0000 Subject: [PATCH 327/505] feat(Algebra/Irreducible): prove irreducibility preserved by `MulEquiv` (#17092) Add two lemmas about Irreducibility preserved by `MulEquiv`. Co-authored-by: jjdishere Co-authored-by: Eric Wieser --- Mathlib/Algebra/Associated/Basic.lean | 23 ++++++++++++++++++++++- 1 file changed, 22 insertions(+), 1 deletion(-) diff --git a/Mathlib/Algebra/Associated/Basic.lean b/Mathlib/Algebra/Associated/Basic.lean index 503074287a0cc..ae0f10085a509 100644 --- a/Mathlib/Algebra/Associated/Basic.lean +++ b/Mathlib/Algebra/Associated/Basic.lean @@ -7,7 +7,7 @@ import Mathlib.Algebra.Group.Even import Mathlib.Algebra.GroupWithZero.Divisibility import Mathlib.Algebra.GroupWithZero.Hom import Mathlib.Algebra.Group.Commute.Units -import Mathlib.Algebra.Group.Units.Hom +import Mathlib.Algebra.Group.Units.Equiv import Mathlib.Order.BoundedOrder import Mathlib.Algebra.Ring.Units @@ -296,6 +296,27 @@ theorem irreducible_mul_iff {a b : M} : · rwa [irreducible_mul_isUnit hb] · rwa [irreducible_isUnit_mul ha] +variable [Monoid N] {F : Type*} [EquivLike F M N] [MulEquivClass F M N] (f : F) + +open MulEquiv + +/-- +Irreducibility is preserved by multiplicative equivalences. +Note that surjective + local hom is not enough. Consider the additive monoids `M = ℕ ⊕ ℕ`, `N = ℕ`, +with a surjective local (additive) hom `f : M →+ N` sending `(m, n)` to `2m + n`. +It is local because the only add unit in `N` is `0`, with preimage `{(0, 0)}` also an add unit. +Then `x = (1, 0)` is irreducible in `M`, but `f x = 2 = 1 + 1` is not irreducible in `N`. +-/ +theorem Irreducible.map {x : M} (h : Irreducible x) : Irreducible (f x) := + ⟨fun g ↦ h.not_unit g.of_map, fun a b g ↦ + let f := MulEquivClass.toMulEquiv f + (h.isUnit_or_isUnit (symm_apply_apply f x ▸ map_mul f.symm a b ▸ congrArg f.symm g)).imp + (·.of_map) (·.of_map)⟩ + +theorem MulEquiv.irreducible_iff (f : F) {a : M} : + Irreducible (f a) ↔ Irreducible a := + ⟨Irreducible.of_map, Irreducible.map f⟩ + end section CommMonoid From f9091e1356447d70b6fa9248e2992380c24b5a9d Mon Sep 17 00:00:00 2001 From: Robin Carlier <57142648+robin-carlier@users.noreply.github.com> Date: Wed, 23 Oct 2024 11:57:25 +0000 Subject: [PATCH 328/505] feat(CategoryTheory/PathCategory): induction principles for morphisms in path categories (#18074) Introduce two induction lemmas to ease proofs on category freely generated on a quiver. --- Mathlib.lean | 3 +- .../CategoryTheory/Bicategory/Coherence.lean | 2 +- Mathlib/CategoryTheory/Category/Quiv.lean | 2 +- .../CategoryTheory/Groupoid/FreeGroupoid.lean | 2 +- .../CategoryTheory/Groupoid/VertexGroup.lean | 2 +- .../Basic.lean} | 49 ++++++++++++++++++- .../PathCategory/MorphismProperty.lean | 43 ++++++++++++++++ Mathlib/Combinatorics/Quiver/Path.lean | 21 ++++++++ 8 files changed, 118 insertions(+), 6 deletions(-) rename Mathlib/CategoryTheory/{PathCategory.lean => PathCategory/Basic.lean} (76%) create mode 100644 Mathlib/CategoryTheory/PathCategory/MorphismProperty.lean diff --git a/Mathlib.lean b/Mathlib.lean index a46fe860dbad2..6986c912408c1 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1840,7 +1840,8 @@ import Mathlib.CategoryTheory.Noetherian import Mathlib.CategoryTheory.Opposites import Mathlib.CategoryTheory.PEmpty import Mathlib.CategoryTheory.PUnit -import Mathlib.CategoryTheory.PathCategory +import Mathlib.CategoryTheory.PathCategory.Basic +import Mathlib.CategoryTheory.PathCategory.MorphismProperty import Mathlib.CategoryTheory.Pi.Basic import Mathlib.CategoryTheory.Preadditive.AdditiveFunctor import Mathlib.CategoryTheory.Preadditive.Basic diff --git a/Mathlib/CategoryTheory/Bicategory/Coherence.lean b/Mathlib/CategoryTheory/Bicategory/Coherence.lean index d22809944d403..b1191ecc4f21c 100644 --- a/Mathlib/CategoryTheory/Bicategory/Coherence.lean +++ b/Mathlib/CategoryTheory/Bicategory/Coherence.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Yuma Mizuno. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yuma Mizuno, Junyan Xu -/ -import Mathlib.CategoryTheory.PathCategory +import Mathlib.CategoryTheory.PathCategory.Basic import Mathlib.CategoryTheory.Functor.FullyFaithful import Mathlib.CategoryTheory.Bicategory.Free import Mathlib.CategoryTheory.Bicategory.LocallyDiscrete diff --git a/Mathlib/CategoryTheory/Category/Quiv.lean b/Mathlib/CategoryTheory/Category/Quiv.lean index be83ad9b038c2..b8ad841e24d51 100644 --- a/Mathlib/CategoryTheory/Category/Quiv.lean +++ b/Mathlib/CategoryTheory/Category/Quiv.lean @@ -5,7 +5,7 @@ Authors: Kim Morrison -/ import Mathlib.CategoryTheory.Adjunction.Basic import Mathlib.CategoryTheory.Category.Cat -import Mathlib.CategoryTheory.PathCategory +import Mathlib.CategoryTheory.PathCategory.Basic /-! # The category of quivers diff --git a/Mathlib/CategoryTheory/Groupoid/FreeGroupoid.lean b/Mathlib/CategoryTheory/Groupoid/FreeGroupoid.lean index 79f597c240e2e..7b6b62d4ce4c8 100644 --- a/Mathlib/CategoryTheory/Groupoid/FreeGroupoid.lean +++ b/Mathlib/CategoryTheory/Groupoid/FreeGroupoid.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémi Bottinelli -/ import Mathlib.CategoryTheory.Groupoid -import Mathlib.CategoryTheory.PathCategory +import Mathlib.CategoryTheory.PathCategory.Basic /-! # Free groupoid on a quiver diff --git a/Mathlib/CategoryTheory/Groupoid/VertexGroup.lean b/Mathlib/CategoryTheory/Groupoid/VertexGroup.lean index 31aaca1e47820..42789d4d9d7d6 100644 --- a/Mathlib/CategoryTheory/Groupoid/VertexGroup.lean +++ b/Mathlib/CategoryTheory/Groupoid/VertexGroup.lean @@ -5,7 +5,7 @@ Authors: Rémi Bottinelli -/ import Mathlib.Algebra.Group.Equiv.Basic import Mathlib.CategoryTheory.Groupoid -import Mathlib.CategoryTheory.PathCategory +import Mathlib.CategoryTheory.PathCategory.Basic import Mathlib.Combinatorics.Quiver.Path /-! diff --git a/Mathlib/CategoryTheory/PathCategory.lean b/Mathlib/CategoryTheory/PathCategory/Basic.lean similarity index 76% rename from Mathlib/CategoryTheory/PathCategory.lean rename to Mathlib/CategoryTheory/PathCategory/Basic.lean index 714ff9a26e3bc..5901349a1ccbd 100644 --- a/Mathlib/CategoryTheory/PathCategory.lean +++ b/Mathlib/CategoryTheory/PathCategory/Basic.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2021 Kim Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Kim Morrison +Authors: Kim Morrison, Robin Carlier -/ import Mathlib.CategoryTheory.EqToHom import Mathlib.CategoryTheory.Quotient @@ -49,6 +49,53 @@ def of : V ⥤q Paths V where obj X := X map f := f.toPath +/-- To prove a property on morphisms of a path category with given source `a`, it suffices to +prove it for the identity and prove that the property is preserved under composition on the right +with length 1 paths. -/ +lemma induction_fixed_source {a : Paths V} (P : ∀ {b : Paths V}, (a ⟶ b) → Prop) + (id : P (𝟙 a)) + (comp : ∀ {u v : V} (p : a ⟶ of.obj u) (q : u ⟶ v), P p → P (p ≫ of.map q)) : + ∀ {b : Paths V} (f : a ⟶ b), P f := by + intro _ f + induction f with + | nil => exact id + | cons _ w h => exact comp _ w h + +/-- To prove a property on morphisms of a path category with given target `b`, it suffices to prove +it for the identity and prove that the property is preserved under composition on the left +with length 1 paths. -/ +lemma induction_fixed_target {b : Paths V} (P : ∀ {a : Paths V}, (a ⟶ b) → Prop) + (id : P (𝟙 b)) + (comp : ∀ {u v : V} (p : of.obj v ⟶ b) (q : u ⟶ v), P p → P (of.map q ≫ p)) : + ∀ {a : Paths V} (f : a ⟶ b), P f := by + intro a f + generalize h : f.length = k + induction k generalizing f a with + | zero => cases f with + | nil => exact id + | cons _ _ => simp at h + | succ k h' => + obtain ⟨c, f, q, hq, rfl⟩ := f.eq_toPath_comp_of_length_eq_succ h + exact comp _ _ (h' _ hq) + +/-- To prove a property on morphisms of a path category, it suffices to prove it for the identity +and prove that the property is preserved under composition on the right with length 1 paths. -/ +lemma induction (P : ∀ {a b : Paths V}, (a ⟶ b) → Prop) + (id : ∀ {v : V}, P (𝟙 (of.obj v))) + (comp : ∀ {u v w : V} (p : of.obj u ⟶ of.obj v) (q : v ⟶ w), P p → P (p ≫ of.map q)) : + ∀ {a b : Paths V} (f : a ⟶ b), P f := + fun {_} ↦ induction_fixed_source _ id comp + +/-- To prove a property on morphisms of a path category, it suffices to prove it for the identity +and prove that the property is preserved under composition on the left with length 1 paths. -/ +lemma induction' (P : ∀ {a b : Paths V}, (a ⟶ b) → Prop) + (id : ∀ {v : V}, P (𝟙 (of.obj v))) + (comp : ∀ {u v w : V} (p : u ⟶ v) (q : of.obj v ⟶ of.obj w), P q → P (of.map p ≫ q)) : + ∀ {a b : Paths V} (f : a ⟶ b), P f := by + intro a b + revert a + exact induction_fixed_target (P := fun f ↦ P f) id (fun _ _ ↦ comp _ _) + attribute [local ext (iff := false)] Functor.ext /-- Any prefunctor from `V` lifts to a functor from `paths V` -/ diff --git a/Mathlib/CategoryTheory/PathCategory/MorphismProperty.lean b/Mathlib/CategoryTheory/PathCategory/MorphismProperty.lean new file mode 100644 index 0000000000000..91b603f5da11d --- /dev/null +++ b/Mathlib/CategoryTheory/PathCategory/MorphismProperty.lean @@ -0,0 +1,43 @@ +/- +Copyright (c) 2024 Robin Carlier. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Robin Carlier +-/ +import Mathlib.CategoryTheory.PathCategory.Basic +import Mathlib.CategoryTheory.MorphismProperty.Basic + +/-! +# Properties of morphisms in a path category. + +We provide a formulation of induction principles for morphisms in a path category in terms of +`MorphismProperty`. This file is separate from `CategoryTheory.PathCategory.Basic` in order to +reduce transitive imports. -/ + + +universe v₁ u₁ + +namespace CategoryTheory.Paths + +variable (V : Type u₁) [Quiver.{v₁ + 1} V] + +/-- A reformulation of `CategoryTheory.Paths.induction` in terms of `MorphismProperty`. -/ +lemma morphismProperty_eq_top + (P : MorphismProperty (Paths V)) + (id : ∀ {v : V}, P (𝟙 (of.obj v))) + (comp : ∀ {u v w : V} (p : of.obj u ⟶ of.obj v) (q : v ⟶ w), P p → P (p ≫ of.map q)) : + P = ⊤ := by + ext; constructor + · simp + · exact fun _ ↦ induction (fun f ↦ P f) id comp _ + +/-- A reformulation of `CategoryTheory.Paths.induction'` in terms of `MorphismProperty`. -/ +lemma morphismProperty_eq_top' + (P : MorphismProperty (Paths V)) + (id : ∀ {v : V}, P (𝟙 (of.obj v))) + (comp : ∀ {u v w : V} (p : u ⟶ v) (q : of.obj v ⟶ of.obj w), P q → P (of.map p ≫ q)) : + P = ⊤ := by + ext; constructor + · simp + · exact fun _ ↦ induction' (fun f ↦ P f) id comp _ + +end CategoryTheory.Paths diff --git a/Mathlib/Combinatorics/Quiver/Path.lean b/Mathlib/Combinatorics/Quiver/Path.lean index d87402e7fb551..39dbcf18bda83 100644 --- a/Mathlib/Combinatorics/Quiver/Path.lean +++ b/Mathlib/Combinatorics/Quiver/Path.lean @@ -72,6 +72,11 @@ theorem eq_of_length_zero (p : Path a b) (hzero : p.length = 0) : a = b := by · rfl · cases Nat.succ_ne_zero _ hzero +theorem eq_nil_of_length_zero (p : Path a a) (hzero : p.length = 0) : p = nil := by + cases p + · rfl + · simp at hzero + /-- Composition of paths. -/ def comp {a b : V} : ∀ {c}, Path a b → Path b c → Path a c | _, p, nil => p @@ -139,6 +144,22 @@ theorem comp_inj_left {p₁ p₂ : Path a b} {q : Path b c} : p₁.comp q = p₂ theorem comp_inj_right {p : Path a b} {q₁ q₂ : Path b c} : p.comp q₁ = p.comp q₂ ↔ q₁ = q₂ := p.comp_injective_right.eq_iff +lemma eq_toPath_comp_of_length_eq_succ (p : Path a b) {n : ℕ} + (hp : p.length = n + 1) : + ∃ (c : V) (f : a ⟶ c) (q : Quiver.Path c b) (_ : q.length = n), + p = f.toPath.comp q := by + induction p generalizing n with + | nil => simp at hp + | @cons c d p q h => + cases n + · rw [length_cons, Nat.zero_add, Nat.add_left_eq_self] at hp + obtain rfl := eq_of_length_zero p hp + obtain rfl := eq_nil_of_length_zero p hp + exact ⟨d, q, nil, rfl, rfl⟩ + · rw [length_cons, Nat.add_right_cancel_iff] at hp + obtain ⟨x, q'', p'', hl, rfl⟩ := h hp + exact ⟨x, q'', p''.cons q, by simpa, rfl⟩ + /-- Turn a path into a list. The list contains `a` at its head, but not `b` a priori. -/ @[simp] def toList : ∀ {b : V}, Path a b → List V From e98856ac3d770804555e9e8c5e96a7512429f10f Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Wed, 23 Oct 2024 12:15:15 +0000 Subject: [PATCH 329/505] refactor: redefine `Module.annihilator` using `RingHom.ker` (#18094) The original definition `LinearMap.ker (LinearMap.lsmul R M)` doesn't work in the noncommutative setting, because `lsmul` isn't `R`-linear if `R` is noncommutative, and also `LinearMap.ker` isn't obviously a two-sided ideal. We also move `Module.annihilator` from Ideal/Operations to Ideal/Maps (where `RingHom.ker` is defined). --- Mathlib/RingTheory/Ideal/Colon.lean | 2 +- Mathlib/RingTheory/Ideal/Maps.lean | 109 +++++++++++++++++++++++ Mathlib/RingTheory/Ideal/Operations.lean | 84 ----------------- 3 files changed, 110 insertions(+), 85 deletions(-) diff --git a/Mathlib/RingTheory/Ideal/Colon.lean b/Mathlib/RingTheory/Ideal/Colon.lean index 903322d97d3f7..9d4dd89b510fc 100644 --- a/Mathlib/RingTheory/Ideal/Colon.lean +++ b/Mathlib/RingTheory/Ideal/Colon.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau -/ import Mathlib.LinearAlgebra.Quotient -import Mathlib.RingTheory.Ideal.Operations +import Mathlib.RingTheory.Ideal.Maps /-! # The colon ideal diff --git a/Mathlib/RingTheory/Ideal/Maps.lean b/Mathlib/RingTheory/Ideal/Maps.lean index 80ddaa2813840..84caddfcb5d16 100644 --- a/Mathlib/RingTheory/Ideal/Maps.lean +++ b/Mathlib/RingTheory/Ideal/Maps.lean @@ -7,6 +7,9 @@ import Mathlib.RingTheory.Ideal.Operations /-! # Maps on modules and ideals + +Main definitions include `Ideal.map`, `Ideal.comap`, `RingHom.ker`, `Module.annihilator` +and `Submodule.annihilator`. -/ assert_not_exists Basis -- See `RingTheory.Ideal.Basis` @@ -637,6 +640,112 @@ theorem ker_isMaximal_of_surjective {R K F : Type*} [Ring R] [Field K] end RingHom +section annihilator + +section Semiring + +variable {R M M' : Type*} +variable [Semiring R] [AddCommMonoid M] [Module R M] [AddCommMonoid M'] [Module R M'] + +variable (R M) in +/-- `Module.annihilator R M` is the ideal of all elements `r : R` such that `r • M = 0`. -/ +def Module.annihilator : Ideal R := RingHom.ker (Module.toAddMonoidEnd R M) + +theorem Module.mem_annihilator {r} : r ∈ Module.annihilator R M ↔ ∀ m : M, r • m = 0 := + ⟨fun h ↦ (congr($h ·)), (AddMonoidHom.ext ·)⟩ + +theorem LinearMap.annihilator_le_of_injective (f : M →ₗ[R] M') (hf : Function.Injective f) : + Module.annihilator R M' ≤ Module.annihilator R M := fun x h ↦ by + rw [Module.mem_annihilator] at h ⊢; exact fun m ↦ hf (by rw [map_smul, h, f.map_zero]) + +theorem LinearMap.annihilator_le_of_surjective (f : M →ₗ[R] M') + (hf : Function.Surjective f) : Module.annihilator R M ≤ Module.annihilator R M' := fun x h ↦ by + rw [Module.mem_annihilator] at h ⊢ + intro m; obtain ⟨m, rfl⟩ := hf m + rw [← map_smul, h, f.map_zero] + +theorem LinearEquiv.annihilator_eq (e : M ≃ₗ[R] M') : + Module.annihilator R M = Module.annihilator R M' := + (e.annihilator_le_of_surjective e.surjective).antisymm (e.annihilator_le_of_injective e.injective) + +namespace Submodule + +/-- `N.annihilator` is the ideal of all elements `r : R` such that `r • N = 0`. -/ +abbrev annihilator (N : Submodule R M) : Ideal R := + Module.annihilator R N + +theorem annihilator_top : (⊤ : Submodule R M).annihilator = Module.annihilator R M := + topEquiv.annihilator_eq + +variable {I J : Ideal R} {N P : Submodule R M} + +theorem mem_annihilator {r} : r ∈ N.annihilator ↔ ∀ n ∈ N, r • n = (0 : M) := by + simp_rw [annihilator, Module.mem_annihilator, Subtype.forall, Subtype.ext_iff]; rfl + +theorem annihilator_bot : (⊥ : Submodule R M).annihilator = ⊤ := + top_le_iff.mp fun _ _ ↦ mem_annihilator.mpr fun _ ↦ by rintro rfl; rw [smul_zero] + +theorem annihilator_eq_top_iff : N.annihilator = ⊤ ↔ N = ⊥ := + ⟨fun H ↦ + eq_bot_iff.2 fun (n : M) hn => + (mem_bot R).2 <| one_smul R n ▸ mem_annihilator.1 ((Ideal.eq_top_iff_one _).1 H) n hn, + fun H ↦ H.symm ▸ annihilator_bot⟩ + +theorem annihilator_mono (h : N ≤ P) : P.annihilator ≤ N.annihilator := fun _ hrp => + mem_annihilator.2 fun n hn => mem_annihilator.1 hrp n <| h hn + +theorem annihilator_iSup (ι : Sort w) (f : ι → Submodule R M) : + annihilator (⨆ i, f i) = ⨅ i, annihilator (f i) := + le_antisymm (le_iInf fun _ => annihilator_mono <| le_iSup _ _) fun r H => + mem_annihilator.2 fun n hn ↦ iSup_induction f (C := (r • · = 0)) hn + (fun i ↦ mem_annihilator.1 <| (mem_iInf _).mp H i) (smul_zero _) + fun m₁ m₂ h₁ h₂ ↦ by simp_rw [smul_add, h₁, h₂, add_zero] + +end Submodule + +end Semiring + +namespace Submodule + +variable {R M : Type*} [CommSemiring R] [AddCommMonoid M] [Module R M] {N : Submodule R M} + +theorem mem_annihilator' {r} : r ∈ N.annihilator ↔ N ≤ comap (r • (LinearMap.id : M →ₗ[R] M)) ⊥ := + mem_annihilator.trans ⟨fun H n hn => (mem_bot R).2 <| H n hn, fun H _ hn => (mem_bot R).1 <| H hn⟩ + +theorem mem_annihilator_span (s : Set M) (r : R) : + r ∈ (Submodule.span R s).annihilator ↔ ∀ n : s, r • (n : M) = 0 := by + rw [Submodule.mem_annihilator] + constructor + · intro h n + exact h _ (Submodule.subset_span n.prop) + · intro h n hn + refine Submodule.span_induction ?_ ?_ ?_ ?_ hn + · intro x hx + exact h ⟨x, hx⟩ + · exact smul_zero _ + · intro x y _ _ hx hy + rw [smul_add, hx, hy, zero_add] + · intro a x _ hx + rw [smul_comm, hx, smul_zero] + +theorem mem_annihilator_span_singleton (g : M) (r : R) : + r ∈ (Submodule.span R ({g} : Set M)).annihilator ↔ r • g = 0 := by simp [mem_annihilator_span] + +@[simp] +theorem annihilator_smul (N : Submodule R M) : annihilator N • N = ⊥ := + eq_bot_iff.2 (smul_le.2 fun _ => mem_annihilator.1) + +@[simp] +theorem annihilator_mul (I : Ideal R) : annihilator I * I = ⊥ := + annihilator_smul I + +@[simp] +theorem mul_annihilator (I : Ideal R) : I * annihilator I = ⊥ := by rw [mul_comm, annihilator_mul] + +end Submodule + +end annihilator + namespace Ideal variable {R : Type*} {S : Type*} {F : Type*} diff --git a/Mathlib/RingTheory/Ideal/Operations.lean b/Mathlib/RingTheory/Ideal/Operations.lean index 8a4763f3a510b..fdc097725c421 100644 --- a/Mathlib/RingTheory/Ideal/Operations.lean +++ b/Mathlib/RingTheory/Ideal/Operations.lean @@ -38,81 +38,8 @@ apply. -/ protected theorem _root_.Ideal.smul_eq_mul (I J : Ideal R) : I • J = I * J := rfl -variable (R M) in -/-- `Module.annihilator R M` is the ideal of all elements `r : R` such that `r • M = 0`. -/ -def _root_.Module.annihilator : Ideal R := LinearMap.ker (LinearMap.lsmul R M) - -theorem _root_.Module.mem_annihilator {r} : r ∈ Module.annihilator R M ↔ ∀ m : M, r • m = 0 := - ⟨fun h ↦ (congr($h ·)), (LinearMap.ext ·)⟩ - -theorem _root_.LinearMap.annihilator_le_of_injective (f : M →ₗ[R] M') (hf : Function.Injective f) : - Module.annihilator R M' ≤ Module.annihilator R M := fun x h ↦ by - rw [Module.mem_annihilator] at h ⊢; exact fun m ↦ hf (by rw [map_smul, h, f.map_zero]) - -theorem _root_.LinearMap.annihilator_le_of_surjective (f : M →ₗ[R] M') - (hf : Function.Surjective f) : Module.annihilator R M ≤ Module.annihilator R M' := fun x h ↦ by - rw [Module.mem_annihilator] at h ⊢ - intro m; obtain ⟨m, rfl⟩ := hf m - rw [← map_smul, h, f.map_zero] - -theorem _root_.LinearEquiv.annihilator_eq (e : M ≃ₗ[R] M') : - Module.annihilator R M = Module.annihilator R M' := - (e.annihilator_le_of_surjective e.surjective).antisymm (e.annihilator_le_of_injective e.injective) - -/-- `N.annihilator` is the ideal of all elements `r : R` such that `r • N = 0`. -/ -abbrev annihilator (N : Submodule R M) : Ideal R := - Module.annihilator R N - -theorem annihilator_top : (⊤ : Submodule R M).annihilator = Module.annihilator R M := - topEquiv.annihilator_eq - variable {I J : Ideal R} {N P : Submodule R M} -theorem mem_annihilator {r} : r ∈ N.annihilator ↔ ∀ n ∈ N, r • n = (0 : M) := by - simp_rw [annihilator, Module.mem_annihilator, Subtype.forall, Subtype.ext_iff]; rfl - -theorem mem_annihilator' {r} : r ∈ N.annihilator ↔ N ≤ comap (r • (LinearMap.id : M →ₗ[R] M)) ⊥ := - mem_annihilator.trans ⟨fun H n hn => (mem_bot R).2 <| H n hn, fun H _ hn => (mem_bot R).1 <| H hn⟩ - -theorem mem_annihilator_span (s : Set M) (r : R) : - r ∈ (Submodule.span R s).annihilator ↔ ∀ n : s, r • (n : M) = 0 := by - rw [Submodule.mem_annihilator] - constructor - · intro h n - exact h _ (Submodule.subset_span n.prop) - · intro h n hn - refine Submodule.span_induction ?_ ?_ ?_ ?_ hn - · intro x hx - exact h ⟨x, hx⟩ - · exact smul_zero _ - · intro x y _ _ hx hy - rw [smul_add, hx, hy, zero_add] - · intro a x _ hx - rw [smul_comm, hx, smul_zero] - -theorem mem_annihilator_span_singleton (g : M) (r : R) : - r ∈ (Submodule.span R ({g} : Set M)).annihilator ↔ r • g = 0 := by simp [mem_annihilator_span] - -theorem annihilator_bot : (⊥ : Submodule R M).annihilator = ⊤ := - (Ideal.eq_top_iff_one _).2 <| mem_annihilator'.2 bot_le - -theorem annihilator_eq_top_iff : N.annihilator = ⊤ ↔ N = ⊥ := - ⟨fun H => - eq_bot_iff.2 fun (n : M) hn => - (mem_bot R).2 <| one_smul R n ▸ mem_annihilator.1 ((Ideal.eq_top_iff_one _).1 H) n hn, - fun H => H.symm ▸ annihilator_bot⟩ - -theorem annihilator_mono (h : N ≤ P) : P.annihilator ≤ N.annihilator := fun _ hrp => - mem_annihilator.2 fun n hn => mem_annihilator.1 hrp n <| h hn - -theorem annihilator_iSup (ι : Sort w) (f : ι → Submodule R M) : - annihilator (⨆ i, f i) = ⨅ i, annihilator (f i) := - le_antisymm (le_iInf fun _ => annihilator_mono <| le_iSup _ _) fun _ H => - mem_annihilator'.2 <| - iSup_le fun i => - have := (mem_iInf _).1 H i - mem_annihilator'.1 this - theorem smul_mem_smul {r} {n} (hr : r ∈ I) (hn : n ∈ N) : r • n ∈ I • N := apply_mem_map₂ _ hr hn @@ -177,17 +104,6 @@ theorem map_le_smul_top (I : Ideal R) (f : R →ₗ[R] M) : rw [← mul_one y, ← smul_eq_mul, f.map_smul] exact smul_mem_smul hy mem_top -@[simp] -theorem annihilator_smul (N : Submodule R M) : annihilator N • N = ⊥ := - eq_bot_iff.2 (smul_le.2 fun _ => mem_annihilator.1) - -@[simp] -theorem annihilator_mul (I : Ideal R) : annihilator I * I = ⊥ := - annihilator_smul I - -@[simp] -theorem mul_annihilator (I : Ideal R) : I * annihilator I = ⊥ := by rw [mul_comm, annihilator_mul] - variable (I J N P) @[simp] From 4e0002f5427da7ca2c65f5311c17c4a86da08273 Mon Sep 17 00:00:00 2001 From: Jz Pan Date: Wed, 23 Oct 2024 12:40:12 +0000 Subject: [PATCH 330/505] doc(NumberTheory.NumberField.House): fix the name of Hua Loo-Keng (#18109) It's L.-K. Hua not H. L. Keng --- Mathlib/NumberTheory/NumberField/House.lean | 2 +- docs/references.bib | 16 ++++++++-------- 2 files changed, 9 insertions(+), 9 deletions(-) diff --git a/Mathlib/NumberTheory/NumberField/House.lean b/Mathlib/NumberTheory/NumberField/House.lean index 18b9930964456..f5cd7772f5a6a 100644 --- a/Mathlib/NumberTheory/NumberField/House.lean +++ b/Mathlib/NumberTheory/NumberField/House.lean @@ -15,7 +15,7 @@ the largest of the modulus of its conjugates. ## References * [D. Marcus, *Number Fields*][marcus1977number] -* [Keng, H. L, *Introduction to number theory*][keng1982house] +* [Hua, L.-K., *Introduction to number theory*][hua1982house] ## Tagshouse number field, algebraic number, house diff --git a/docs/references.bib b/docs/references.bib index f7bd7c609f40e..ecdb8ed7428cd 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -1912,6 +1912,14 @@ @Misc{ howard author = {Howard, Peter} } +@Book{ hua1982house, + author = {Hua, Loo-Keng}, + publisher = {Springer}, + title = {Introduction to Number Theory}, + year = {1982}, + pages = {489} +} + @Book{ HubbardWest-ode, author = {John H. Hubbard and Beverly H. West}, title = {Differential Equations: A Dynamical Systems Approach}, @@ -2242,14 +2250,6 @@ @Article{ kelleyVaught1953 doi = {10.2307/1990847} } -@Book{ keng1982house, - author = {Keng, Hua Loo}, - publisher = {Springer}, - title = {Introduction to Number Theory}, - year = {1982}, - pages = {489} -} - @Article{ kleiman1979, author = {Kleiman, Steven Lawrence}, title = {Misconceptions about {$K_X$}}, From 4dac3e5656e4ba39002781c24e26d3847fb10a63 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Wed, 23 Oct 2024 12:50:44 +0000 Subject: [PATCH 331/505] chore: re-add `@[simp]` attributes lost in the port (#18076) This PR re-adds the `@[simp]` attributes that had a porting note that they could be proved by `@[simp]`, but where that does not appear to be the case currently. Of course, the `simp` set may have changed sufficiently to require un`@[simp]`ing some of these anyway, but we should add a more up to date comment in that case. --- Mathlib/Analysis/RCLike/Basic.lean | 2 +- Mathlib/CategoryTheory/Comma/Over.lean | 4 ++-- Mathlib/CategoryTheory/Conj.lean | 8 +++---- .../Limits/Shapes/BinaryProducts.lean | 7 +++--- .../Limits/Shapes/Biproducts.lean | 3 +-- .../Limits/Shapes/WidePullbacks.lean | 1 - Mathlib/Data/Finmap.lean | 7 +++++- Mathlib/Data/List/Basic.lean | 24 +++++++------------ Mathlib/Data/List/Sigma.lean | 2 +- Mathlib/Data/NNRat/Defs.lean | 4 ++-- Mathlib/Data/Num/Lemmas.lean | 2 +- Mathlib/Data/Option/NAry.lean | 2 +- 12 files changed, 30 insertions(+), 36 deletions(-) diff --git a/Mathlib/Analysis/RCLike/Basic.lean b/Mathlib/Analysis/RCLike/Basic.lean index ad002383d932e..6df3640fd74f1 100644 --- a/Mathlib/Analysis/RCLike/Basic.lean +++ b/Mathlib/Analysis/RCLike/Basic.lean @@ -299,7 +299,7 @@ theorem conj_nat_cast (n : ℕ) : conj (n : K) = n := map_natCast _ _ theorem conj_ofNat (n : ℕ) [n.AtLeastTwo] : conj (no_index (OfNat.ofNat n : K)) = OfNat.ofNat n := map_ofNat _ _ -@[rclike_simps] -- Porting note (#10618): was a `simp` but `simp` can prove it +@[rclike_simps, simp] theorem conj_neg_I : conj (-I) = (I : K) := by rw [map_neg, conj_I, neg_neg] theorem conj_eq_re_sub_im (z : K) : conj z = re z - im z * I := diff --git a/Mathlib/CategoryTheory/Comma/Over.lean b/Mathlib/CategoryTheory/Comma/Over.lean index 12c572b50df5c..3b9a32d2c8eee 100644 --- a/Mathlib/CategoryTheory/Comma/Over.lean +++ b/Mathlib/CategoryTheory/Comma/Over.lean @@ -59,7 +59,7 @@ theorem OverMorphism.ext {X : T} {U V : Over X} {f g : U ⟶ V} (h : f.left = g. congr simp only [eq_iff_true_of_subsingleton] --- @[simp] : Porting note (#10618): simp can prove this +@[simp] theorem over_right (U : Over X) : U.right = ⟨⟨⟩⟩ := by simp only @[simp] @@ -369,7 +369,7 @@ theorem UnderMorphism.ext {X : T} {U V : Under X} {f g : U ⟶ V} (h : f.right = let ⟨_,b,_⟩ := f; let ⟨_,e,_⟩ := g congr; simp only [eq_iff_true_of_subsingleton] --- @[simp] Porting note (#10618): simp can prove this +@[simp] theorem under_left (U : Under X) : U.left = ⟨⟨⟩⟩ := by simp only @[simp] diff --git a/Mathlib/CategoryTheory/Conj.lean b/Mathlib/CategoryTheory/Conj.lean index 5246270f857c8..e0598514c3486 100644 --- a/Mathlib/CategoryTheory/Conj.lean +++ b/Mathlib/CategoryTheory/Conj.lean @@ -63,7 +63,7 @@ theorem symm_self_conj (f : End X) : α.symm.conj (α.conj f) = f := by theorem self_symm_conj (f : End Y) : α.conj (α.symm.conj f) = f := α.symm.symm_self_conj f -/- Porting note (#10618): removed `@[simp]`; simp can prove this -/ +@[simp] theorem conj_pow (f : End X) (n : ℕ) : α.conj (f ^ n) = α.conj f ^ n := α.conj.toMonoidHom.map_pow f n @@ -83,7 +83,7 @@ theorem trans_conjAut {Z : C} (β : Y ≅ Z) (f : Aut X) : (α ≪≫ β).conjAut f = β.conjAut (α.conjAut f) := by simp only [conjAut_apply, Iso.trans_symm, Iso.trans_assoc] -/- Porting note (#10618): removed `@[simp]`; simp can prove this -/ +@[simp] theorem conjAut_mul (f g : Aut X) : α.conjAut (f * g) = α.conjAut f * α.conjAut g := map_mul α.conjAut f g @@ -91,11 +91,11 @@ theorem conjAut_mul (f g : Aut X) : α.conjAut (f * g) = α.conjAut f * α.conjA theorem conjAut_trans (f g : Aut X) : α.conjAut (f ≪≫ g) = α.conjAut f ≪≫ α.conjAut g := conjAut_mul α g f -/- Porting note (#10618): removed `@[simp]`; simp can prove this -/ +@[simp] theorem conjAut_pow (f : Aut X) (n : ℕ) : α.conjAut (f ^ n) = α.conjAut f ^ n := map_pow α.conjAut f n -/- Porting note (#10618): removed `@[simp]`; simp can prove this -/ +@[simp] theorem conjAut_zpow (f : Aut X) (n : ℤ) : α.conjAut (f ^ n) = α.conjAut f ^ n := map_zpow α.conjAut f n diff --git a/Mathlib/CategoryTheory/Limits/Shapes/BinaryProducts.lean b/Mathlib/CategoryTheory/Limits/Shapes/BinaryProducts.lean index 841fb4ae1579b..3cf40588ccb09 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/BinaryProducts.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/BinaryProducts.lean @@ -559,7 +559,6 @@ noncomputable abbrev coprod.desc {W X Y : C} [HasBinaryCoproduct X Y] noncomputable abbrev codiag (X : C) [HasBinaryCoproduct X X] : X ⨿ X ⟶ X := coprod.desc (𝟙 _) (𝟙 _) --- Porting note (#10618): simp removes as simp can prove this @[reassoc] theorem prod.lift_fst {W X Y : C} [HasBinaryProduct X Y] (f : W ⟶ X) (g : W ⟶ Y) : prod.lift f g ≫ prod.fst = f := @@ -709,15 +708,15 @@ instance prod.map_mono {C : Type*} [Category C] {W X Y Z : C} (f : W ⟶ Y) (g : · rw [← cancel_mono g] simpa using congr_arg (fun f => f ≫ prod.snd) h⟩ -@[reassoc] -- Porting note (#10618): simp can prove these +@[reassoc] theorem prod.diag_map {X Y : C} (f : X ⟶ Y) [HasBinaryProduct X X] [HasBinaryProduct Y Y] : diag X ≫ prod.map f f = f ≫ diag Y := by simp -@[reassoc] -- Porting note (#10618): simp can prove these +@[reassoc] theorem prod.diag_map_fst_snd {X Y : C} [HasBinaryProduct X Y] [HasBinaryProduct (X ⨯ Y) (X ⨯ Y)] : diag (X ⨯ Y) ≫ prod.map prod.fst prod.snd = 𝟙 (X ⨯ Y) := by simp -@[reassoc] -- Porting note (#10618): simp can prove these +@[reassoc] theorem prod.diag_map_fst_snd_comp [HasLimitsOfShape (Discrete WalkingPair) C] {X X' Y Y' : C} (g : X ⟶ Y) (g' : X' ⟶ Y') : diag (X ⨯ X') ≫ prod.map (prod.fst ≫ g) (prod.snd ≫ g') = prod.map g g' := by simp diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean b/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean index 5bf06db5fdc26..89b2d0b933bb6 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean @@ -250,8 +250,7 @@ structure IsBilimit {F : J → C} (B : Bicone F) where attribute [inherit_doc IsBilimit] IsBilimit.isLimit IsBilimit.isColimit --- Porting note (#10618): simp can prove this, linter doesn't notice it is removed -attribute [-simp, nolint simpNF] IsBilimit.mk.injEq +attribute [simp] IsBilimit.mk.injEq attribute [local ext] Bicone.IsBilimit diff --git a/Mathlib/CategoryTheory/Limits/Shapes/WidePullbacks.lean b/Mathlib/CategoryTheory/Limits/Shapes/WidePullbacks.lean index bb674ca92f2a0..4670206841758 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/WidePullbacks.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/WidePullbacks.lean @@ -376,7 +376,6 @@ noncomputable abbrev desc {X : C} (f : B ⟶ X) (fs : ∀ j : J, objs j ⟶ X) variable (arrows) variable {X : C} (f : B ⟶ X) (fs : ∀ j : J, objs j ⟶ X) (w : ∀ j, arrows j ≫ fs j = f) --- Porting note (#10618): simp can prove this so removed simp attribute @[reassoc] theorem ι_desc (j : J) : ι arrows j ≫ desc f fs w = fs _ := by simp only [colimit.ι_desc, WidePushoutShape.mkCocone_pt, WidePushoutShape.mkCocone_ι_app] diff --git a/Mathlib/Data/Finmap.lean b/Mathlib/Data/Finmap.lean index 9d6ec0b616a21..0b0ae8a95ac49 100644 --- a/Mathlib/Data/Finmap.lean +++ b/Mathlib/Data/Finmap.lean @@ -513,7 +513,12 @@ theorem lookup_union_left_of_not_in {a} {s₁ s₂ : Finmap β} (h : a ∉ s₂) · rw [lookup_union_left h'] · rw [lookup_union_right h', lookup_eq_none.mpr h, lookup_eq_none.mpr h'] --- @[simp] -- Porting note (#10618): simp can prove this +/-- `simp`-normal form of `mem_lookup_union` -/ +@[simp] +theorem mem_lookup_union' {a} {b : β a} {s₁ s₂ : Finmap β} : + lookup a (s₁ ∪ s₂) = some b ↔ b ∈ lookup a s₁ ∨ a ∉ s₁ ∧ b ∈ lookup a s₂ := + induction_on₂ s₁ s₂ fun _ _ => AList.mem_lookup_union + theorem mem_lookup_union {a} {b : β a} {s₁ s₂ : Finmap β} : b ∈ lookup a (s₁ ∪ s₂) ↔ b ∈ lookup a s₁ ∨ a ∉ s₁ ∧ b ∈ lookup a s₂ := induction_on₂ s₁ s₂ fun _ _ => AList.mem_lookup_union diff --git a/Mathlib/Data/List/Basic.lean b/Mathlib/Data/List/Basic.lean index 74c813e67ec38..efee2daa93ea1 100644 --- a/Mathlib/Data/List/Basic.lean +++ b/Mathlib/Data/List/Basic.lean @@ -261,8 +261,7 @@ theorem reverse_cons' (a : α) (l : List α) : reverse (a :: l) = concat (revers theorem reverse_concat' (l : List α) (a : α) : (l ++ [a]).reverse = a :: l.reverse := by rw [reverse_append]; rfl --- Porting note (#10618): simp can prove this --- @[simp] +@[simp] theorem reverse_singleton (a : α) : reverse [a] = [a] := rfl @@ -308,8 +307,7 @@ theorem getLast_concat' {a : α} (l : List α) : getLast (concat l a) (concat_ne @[simp] theorem getLast_singleton' (a : α) : getLast [a] (cons_ne_nil a []) = a := rfl --- Porting note (#10618): simp can prove this --- @[simp] +@[simp] theorem getLast_cons_cons (a₁ a₂ : α) (l : List α) : getLast (a₁ :: a₂ :: l) (cons_ne_nil _ _) = getLast (a₂ :: l) (cons_ne_nil a₂ l) := rfl @@ -1947,8 +1945,7 @@ theorem map₂Right'_nil_left : map₂Right' f [] bs = (bs.map (f none), []) := theorem map₂Right'_nil_right : map₂Right' f as [] = ([], as) := rfl --- Porting note (#10618): simp can prove this --- @[simp] +@[simp] theorem map₂Right'_nil_cons : map₂Right' f [] (b :: bs) = (f none b :: bs.map (f none), []) := rfl @@ -1975,8 +1972,7 @@ theorem zipLeft'_nil_right : zipLeft' as ([] : List β) = (as.map fun a => (a, n theorem zipLeft'_nil_left : zipLeft' ([] : List α) bs = ([], bs) := rfl --- Porting note (#10618): simp can prove this --- @[simp] +@[simp] theorem zipLeft'_cons_nil : zipLeft' (a :: as) ([] : List β) = ((a, none) :: as.map fun a => (a, none), []) := rfl @@ -2004,8 +2000,7 @@ theorem zipRight'_nil_left : zipRight' ([] : List α) bs = (bs.map fun b => (non theorem zipRight'_nil_right : zipRight' as ([] : List β) = ([], as) := rfl --- Porting note (#10618): simp can prove this --- @[simp] +@[simp] theorem zipRight'_nil_cons : zipRight' ([] : List α) (b :: bs) = ((none, b) :: bs.map fun b => (none, b), []) := rfl @@ -2060,8 +2055,7 @@ theorem map₂Right_nil_left : map₂Right f [] bs = bs.map (f none) := by cases theorem map₂Right_nil_right : map₂Right f as [] = [] := rfl --- Porting note (#10618): simp can prove this --- @[simp] +@[simp] theorem map₂Right_nil_cons : map₂Right f [] (b :: bs) = f none b :: bs.map (f none) := rfl @@ -2094,8 +2088,7 @@ theorem zipLeft_nil_right : zipLeft as ([] : List β) = as.map fun a => (a, none theorem zipLeft_nil_left : zipLeft ([] : List α) bs = [] := rfl --- Porting note (#10618): simp can prove this --- @[simp] +@[simp] theorem zipLeft_cons_nil : zipLeft (a :: as) ([] : List β) = (a, none) :: as.map fun a => (a, none) := rfl @@ -2132,8 +2125,7 @@ theorem zipRight_nil_left : zipRight ([] : List α) bs = bs.map fun b => (none, theorem zipRight_nil_right : zipRight as ([] : List β) = [] := rfl --- Porting note (#10618): simp can prove this --- @[simp] +@[simp] theorem zipRight_nil_cons : zipRight ([] : List α) (b :: bs) = (none, b) :: bs.map fun b => (none, b) := rfl diff --git a/Mathlib/Data/List/Sigma.lean b/Mathlib/Data/List/Sigma.lean index 7064c0948f068..06fa67f877f9c 100644 --- a/Mathlib/Data/List/Sigma.lean +++ b/Mathlib/Data/List/Sigma.lean @@ -365,7 +365,7 @@ theorem Perm.kreplace {a : α} {b : β a} {l₁ l₂ : List (Sigma β)} (nd : l def kerase (a : α) : List (Sigma β) → List (Sigma β) := eraseP fun s => a = s.1 --- Porting note (#10618): removing @[simp], `simp` can prove it +@[simp] theorem kerase_nil {a} : @kerase _ β _ a [] = [] := rfl diff --git a/Mathlib/Data/NNRat/Defs.lean b/Mathlib/Data/NNRat/Defs.lean index 3f411698b84cd..e0b69c3d877c3 100644 --- a/Mathlib/Data/NNRat/Defs.lean +++ b/Mathlib/Data/NNRat/Defs.lean @@ -147,11 +147,11 @@ theorem coe_eq_zero : (q : ℚ) = 0 ↔ q = 0 := by norm_cast theorem coe_ne_zero : (q : ℚ) ≠ 0 ↔ q ≠ 0 := coe_eq_zero.not -@[norm_cast] -- Porting note (#10618): simp can prove this +@[norm_cast] theorem coe_le_coe : (p : ℚ) ≤ q ↔ p ≤ q := Iff.rfl -@[norm_cast] -- Porting note (#10618): simp can prove this +@[norm_cast] theorem coe_lt_coe : (p : ℚ) < q ↔ p < q := Iff.rfl diff --git a/Mathlib/Data/Num/Lemmas.lean b/Mathlib/Data/Num/Lemmas.lean index 1d1f99cc29e6a..8e33a4752ee9e 100644 --- a/Mathlib/Data/Num/Lemmas.lean +++ b/Mathlib/Data/Num/Lemmas.lean @@ -45,7 +45,7 @@ theorem cast_to_nat [AddMonoidWithOne α] : ∀ n : PosNum, ((n : ℕ) : α) = n | bit0 p => by dsimp; rw [Nat.cast_add, p.cast_to_nat] | bit1 p => by dsimp; rw [Nat.cast_add, Nat.cast_add, Nat.cast_one, p.cast_to_nat] -@[norm_cast] -- @[simp] -- Porting note (#10618): simp can prove this +@[norm_cast] theorem to_nat_to_int (n : PosNum) : ((n : ℕ) : ℤ) = n := cast_to_nat _ diff --git a/Mathlib/Data/Option/NAry.lean b/Mathlib/Data/Option/NAry.lean index 855ae63c32d0a..b6ee71344701b 100644 --- a/Mathlib/Data/Option/NAry.lean +++ b/Mathlib/Data/Option/NAry.lean @@ -45,7 +45,7 @@ theorem map₂_def {α β γ : Type u} (f : α → β → γ) (a : Option α) (b map₂ f a b = f <$> a <*> b := by cases a <;> rfl --- Porting note (#10618): In Lean3, was `@[simp]` but now `simp` can prove it +@[simp] theorem map₂_some_some (f : α → β → γ) (a : α) (b : β) : map₂ f (some a) (some b) = f a b := rfl theorem map₂_coe_coe (f : α → β → γ) (a : α) (b : β) : map₂ f a b = f a b := rfl From 3b90821802963a22611f445f981ef27e280d1335 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Wed, 23 Oct 2024 13:28:02 +0000 Subject: [PATCH 332/505] feat(RingTheory/Flat/FaithfullyFlat): faithful flatness in terms of exactness (#17869) $M$ is faithfully flat if and only if $M$ is flat and $- \otimes M$ preserves and reflects short exact sequences --- .../TensorProduct/RightExactness.lean | 15 +- Mathlib/RingTheory/Flat/FaithfullyFlat.lean | 180 ++++++++++++++++++ 2 files changed, 189 insertions(+), 6 deletions(-) diff --git a/Mathlib/LinearAlgebra/TensorProduct/RightExactness.lean b/Mathlib/LinearAlgebra/TensorProduct/RightExactness.lean index a61f53baf761c..89591b4a9ce54 100644 --- a/Mathlib/LinearAlgebra/TensorProduct/RightExactness.lean +++ b/Mathlib/LinearAlgebra/TensorProduct/RightExactness.lean @@ -157,6 +157,13 @@ theorem LinearMap.rTensor_range : apply rTensor_surjective rw [← range_eq_top, range_rangeRestrict] +lemma LinearMap.rTensor_exact_iff_lTensor_exact : + Function.Exact (f.rTensor Q) (g.rTensor Q) ↔ + Function.Exact (f.lTensor Q) (g.lTensor Q) := + Function.Exact.iff_of_ladder_linearEquiv (e₁ := TensorProduct.comm _ _ _) + (e₂ := TensorProduct.comm _ _ _) (e₃ := TensorProduct.comm _ _ _) + (by ext; simp) (by ext; simp) + end Semiring variable {R M N P : Type*} [CommRing R] @@ -370,12 +377,8 @@ noncomputable def rTensor.equiv : include hfg hg in /-- Tensoring an exact pair on the right gives an exact pair -/ theorem rTensor_exact : Exact (rTensor Q f) (rTensor Q g) := by - rw [exact_iff, ← Submodule.ker_mkQ (p := range (rTensor Q f)), - ← rTensor.inverse_comp_rTensor Q hfg hg] - apply symm - apply ker_comp_of_ker_eq_bot - rw [ker_eq_bot] - exact (rTensor.equiv Q hfg hg).symm.injective + rw [rTensor_exact_iff_lTensor_exact] + exact lTensor_exact Q hfg hg /-- Right-exactness of tensor product (`rTensor`) -/ lemma rTensor_mkQ (N : Submodule R M) : diff --git a/Mathlib/RingTheory/Flat/FaithfullyFlat.lean b/Mathlib/RingTheory/Flat/FaithfullyFlat.lean index 17c82bb7c375d..8f31c2d917255 100644 --- a/Mathlib/RingTheory/Flat/FaithfullyFlat.lean +++ b/Mathlib/RingTheory/Flat/FaithfullyFlat.lean @@ -25,6 +25,12 @@ A module `M` over a commutative ring `R` is *faithfully flat* if it is flat and is flat and tensoring with `M` is faithful, i.e. `N ≠ 0` implies `N ⊗ M ≠ 0`. - `Module.FaithfullyFlat.iff_flat_and_lTensor_faithful`: an `R`-module `M` is faithfully flat iff it is flat and tensoring with `M` is faithful, i.e. `N ≠ 0` implies `M ⊗ N ≠ 0`. +- `Module.FaithfullyFlat.iff_exact_iff_rTensor_exact`: an `R`-module `M` is faithfully flat iff + tensoring with `M` preserves and reflects exact sequences, i.e. the sequence `N₁ → N₂ → N₃` is + exact *iff* the sequence `N₁ ⊗ M → N₂ ⊗ M → N₃ ⊗ M` is exact. +- `Module.FaithfullyFlat.iff_exact_iff_lTensor_exact`: an `R`-module `M` is faithfully flat iff + tensoring with `M` preserves and reflects exact sequences, i.e. the sequence `N₁ → N₂ → N₃` is + exact *iff* the sequence `M ⊗ N₁ → M ⊗ N₂ → M ⊗ N₃` is exact. - `Module.FaithfullyFlat.self`: the `R`-module `R` is faithfully flat. @@ -163,6 +169,180 @@ lemma iff_flat_and_lTensor_reflects_triviality : end faithful +section exact + +/-! +### Faithfully flat modules and exact sequences + +In this section we prove that an `R`-module `M` is faithfully flat iff tensoring with `M` +preserves and reflects exact sequences. + +Let `N₁ -l₁₂-> N₂ -l₂₃-> N₃` be two linear maps. +- We first show that if `N₁ ⊗ M -> N₂ ⊗ M -> N₃ ⊗ M` is exact, then `N₁ -l₁₂-> N₂ -l₂₃-> N₃` is a + complex, i.e. `range l₁₂ ≤ ker l₂₃`. + This is `range_le_ker_of_exact_rTensor`. +- Then in `rTensor_reflects_exact`, we show `ker l₂₃ = range l₁₂` by considering the cohomology + `ker l₂₃ ⧸ range l₁₂`. +This shows that when `M` is faithfully flat, `- ⊗ M` reflects exact sequences. For details, see +comments in the proof. Since `M` is flat, `- ⊗ M` preserves exact sequences. + +On the other hand, if `- ⊗ M` preserves and reflects exact sequences, then `M` is faithfully flat. +- `M` is flat because `- ⊗ M` preserves exact sequences. +- We need to show that if `N ⊗ M = 0` then `N = 0`. Consider the sequence `N -0-> N -0-> 0`. After + tensoring with `M`, we get `N ⊗ M -0-> N ⊗ M -0-> 0` which is exact because `N ⊗ M = 0`. + Since `- ⊗ M` reflects exact sequences, `N = 0`. +-/ + +section arbitrary_universe + +variable {N1 : Type*} [AddCommGroup N1] [Module R N1] +variable {N2 : Type*} [AddCommGroup N2] [Module R N2] +variable {N3 : Type*} [AddCommGroup N3] [Module R N3] +variable (l12 : N1 →ₗ[R] N2) (l23 : N2 →ₗ[R] N3) + +/-- +If `M` is faithfully flat, then exactness of `N₁ ⊗ M -> N₂ ⊗ M -> N₃ ⊗ M` implies that the +composition `N₁ -> N₂ -> N₃` is `0`. + +Implementation detail, please use `rTensor_reflects_exact` instead. +-/ +lemma range_le_ker_of_exact_rTensor [fl : FaithfullyFlat R M] + (ex : Function.Exact (l12.rTensor M) (l23.rTensor M)) : + LinearMap.range l12 ≤ LinearMap.ker l23 := by + -- let `n1 ∈ N1`. We need to show `l23 (l12 n1) = 0`. Suppose this is not the case. + rintro _ ⟨n1, rfl⟩ + rw [LinearMap.mem_ker] + by_contra! hn1 + -- Let `E` be the submodule spanned by `l23 (l12 n1)`. Then because `l23 (l12 n1) ≠ 0`, we have + -- `E ≠ 0`. + let E : Submodule R N3 := Submodule.span R {l23 (l12 n1)} + have hE : Nontrivial E := + ⟨0, ⟨⟨l23 (l12 n1), Submodule.mem_span_singleton_self _⟩, Subtype.coe_ne_coe.1 hn1.symm⟩⟩ + + -- Since `N1 ⊗ M -> N2 ⊗ M -> N3 ⊗ M` is exact, we have `l23 (l12 n1) ⊗ₜ m = 0` for all `m : M`. + have eq1 : ∀ (m : M), l23 (l12 n1) ⊗ₜ[R] m = 0 := fun m ↦ + ex.apply_apply_eq_zero (n1 ⊗ₜ[R] m) + -- Then `E ⊗ M = 0`. Indeed, + have eq0 : (⊤ : Submodule R (E ⊗[R] M)) = ⊥ := by + -- suppose `x ∈ E ⊗ M`. We will show `x = 0`. + ext x + simp only [Submodule.mem_top, Submodule.mem_bot, true_iff] + have mem : x ∈ (⊤ : Submodule R _) := ⟨⟩ + rw [← TensorProduct.span_tmul_eq_top, mem_span_set] at mem + obtain ⟨c, hc, rfl⟩ := mem + choose b a hy using hc + let r : ⦃a : E ⊗[R] M⦄ → a ∈ ↑c.support → R := fun a ha => + Submodule.mem_span_singleton.1 (b ha).2 |>.choose + have hr : ∀ ⦃i : E ⊗[R] M⦄ (hi : i ∈ c.support), b hi = + r hi • ⟨l23 (l12 n1), Submodule.mem_span_singleton_self _⟩ := fun a ha => + Subtype.ext <| Submodule.mem_span_singleton.1 (b ha).2 |>.choose_spec.symm + -- Since `M` is flat and `E -> N1` is injective, we only need to check that x = 0 + -- in `N1 ⊗ M`. We write `x = ∑ μᵢ • (l23 (l12 n1)) ⊗ mᵢ = ∑ μᵢ • 0 = 0` + -- (remember `E = span {l23 (l12 n1)}` and `eq1`) + refine Finset.sum_eq_zero fun i hi => show c i • i = 0 from + (Module.Flat.rTensor_preserves_injective_linearMap (M := M) E.subtype <| + Submodule.injective_subtype E) ?_ + rw [← hy hi, hr hi, smul_tmul, map_smul, LinearMap.rTensor_tmul, Submodule.subtype_apply, eq1, + smul_zero, map_zero] + have : Subsingleton (E ⊗[R] M) := subsingleton_iff_forall_eq 0 |>.2 fun x => + show x ∈ (⊥ : Submodule R _) from eq0 ▸ ⟨⟩ + + -- but `E ⊗ M = 0` implies `E = 0` because `M` is faithfully flat and this is a contradiction. + exact not_subsingleton_iff_nontrivial.2 inferInstance <| fl.rTensor_reflects_triviality R M E + +lemma rTensor_reflects_exact [fl : FaithfullyFlat R M] + (ex : Function.Exact (l12.rTensor M) (l23.rTensor M)) : + Function.Exact l12 l23 := LinearMap.exact_iff.2 <| by + have complex : LinearMap.range l12 ≤ LinearMap.ker l23 := range_le_ker_of_exact_rTensor R M _ _ ex + -- By the previous lemma we have that range l12 ≤ ker l23 and hence the quotient + -- H := ker l23 ⧸ range l12 makes sense. + -- Hence our goal ker l23 = range l12 follows from the claim that H = 0. + let H := LinearMap.ker l23 ⧸ LinearMap.range (Submodule.inclusion complex) + suffices triv_coh : Subsingleton H by + rw [Submodule.subsingleton_quotient_iff_eq_top, Submodule.range_inclusion, + Submodule.comap_subtype_eq_top] at triv_coh + exact le_antisymm triv_coh complex + + -- Since `M` is faithfully flat, we need only to show that `H ⊗ M` is trivial. + suffices Subsingleton (H ⊗[R] M) from rTensor_reflects_triviality R M H + let e : H ⊗[R] M ≃ₗ[R] _ := TensorProduct.quotientTensorEquiv _ _ + -- Note that `H ⊗ M` is isomorphic to `ker l12 ⊗ M ⧸ range ((range l12 ⊗ M) -> (ker l23 ⊗ M))`. + -- So the problem is reduced to proving surjectivity of `range l12 ⊗ M → ker l23 ⊗ M`. + rw [e.toEquiv.subsingleton_congr, Submodule.subsingleton_quotient_iff_eq_top, + LinearMap.range_eq_top] + intro x + induction x using TensorProduct.induction_on with + | zero => exact ⟨0, by simp⟩ + -- let `x ⊗ m` be an element in `ker l23 ⊗ M`, then `x ⊗ m` is in the kernel of `l23 ⊗ 𝟙M`. + -- Since `N1 ⊗ M -l12 ⊗ M-> N2 ⊗ M -l23 ⊗ M-> N3 ⊗ M` is exact, we have that `x ⊗ m` is in + -- the range of `l12 ⊗ 𝟙M`, i.e. `x ⊗ m = (l12 ⊗ 𝟙M) y` for some `y ∈ N1 ⊗ M` as elements of + -- `N2 ⊗ M`. We need to prove that `x ⊗ m = (l12 ⊗ 𝟙M) y` still holds in `(ker l23) ⊗ M`. + -- This is okay because `M` is flat and `ker l23 -> N2` is injective. + | tmul x m => + rcases x with ⟨x, (hx : l23 x = 0)⟩ + have mem : x ⊗ₜ[R] m ∈ LinearMap.ker (l23.rTensor M) := by simp [hx] + rw [LinearMap.exact_iff.1 ex] at mem + obtain ⟨y, hy⟩ := mem + + refine ⟨LinearMap.rTensor M (LinearMap.rangeRestrict _ ∘ₗ LinearMap.rangeRestrict l12) y, + Module.Flat.rTensor_preserves_injective_linearMap (LinearMap.ker l23).subtype + Subtype.val_injective ?_⟩ + simp only [LinearMap.comp_codRestrict, LinearMap.rTensor_tmul, Submodule.coe_subtype, ← hy] + rw [← LinearMap.comp_apply] + erw [← LinearMap.rTensor_comp] + rw [← LinearMap.comp_apply, ← LinearMap.rTensor_comp, LinearMap.comp_assoc, + LinearMap.subtype_comp_codRestrict, ← LinearMap.comp_assoc, Submodule.subtype_comp_inclusion, + LinearMap.subtype_comp_codRestrict] + | add x y hx hy => + obtain ⟨x, rfl⟩ := hx; obtain ⟨y, rfl⟩ := hy + exact ⟨x + y, by simp⟩ + +lemma lTensor_reflects_exact [fl : FaithfullyFlat R M] + (ex : Function.Exact (l12.lTensor M) (l23.lTensor M)) : + Function.Exact l12 l23 := + rTensor_reflects_exact R M _ _ <| ex.of_ladder_linearEquiv_of_exact + (e₁ := TensorProduct.comm _ _ _) (e₂ := TensorProduct.comm _ _ _) + (e₃ := TensorProduct.comm _ _ _) (by ext; rfl) (by ext; rfl) + +end arbitrary_universe + +section fixed_universe + +lemma exact_iff_rTensor_exact [fl : FaithfullyFlat R M] + {N1 : Type max u v} [AddCommGroup N1] [Module R N1] + {N2 : Type max u v} [AddCommGroup N2] [Module R N2] + {N3 : Type max u v} [AddCommGroup N3] [Module R N3] + (l12 : N1 →ₗ[R] N2) (l23 : N2 →ₗ[R] N3) : + Function.Exact l12 l23 ↔ Function.Exact (l12.rTensor M) (l23.rTensor M) := + ⟨fun e => Module.Flat.iff_rTensor_exact.1 fl.toFlat e, + fun ex => rTensor_reflects_exact R M l12 l23 ex⟩ + +lemma iff_exact_iff_rTensor_exact : + FaithfullyFlat R M ↔ + (∀ {N1 : Type max u v} [AddCommGroup N1] [Module R N1] + {N2 : Type max u v} [AddCommGroup N2] [Module R N2] + {N3 : Type max u v} [AddCommGroup N3] [Module R N3] + (l12 : N1 →ₗ[R] N2) (l23 : N2 →ₗ[R] N3), + Function.Exact l12 l23 ↔ Function.Exact (l12.rTensor M) (l23.rTensor M)) := + ⟨fun fl => exact_iff_rTensor_exact R M, fun iff_exact => + iff_flat_and_rTensor_reflects_triviality _ _ |>.2 ⟨Flat.iff_rTensor_exact.2 <| by aesop, + fun N _ _ h => subsingleton_iff_forall_eq 0 |>.2 <| fun y => by + simpa [eq_comm] using (iff_exact (0 : PUnit →ₗ[R] N) (0 : N →ₗ[R] PUnit) |>.2 fun x => by + simpa using Subsingleton.elim _ _) y⟩⟩ + +lemma iff_exact_iff_lTensor_exact : + FaithfullyFlat R M ↔ + (∀ {N1 : Type max u v} [AddCommGroup N1] [Module R N1] + {N2 : Type max u v} [AddCommGroup N2] [Module R N2] + {N3 : Type max u v} [AddCommGroup N3] [Module R N3] + (l12 : N1 →ₗ[R] N2) (l23 : N2 →ₗ[R] N3), + Function.Exact l12 l23 ↔ Function.Exact (l12.lTensor M) (l23.lTensor M)) := by + simp only [iff_exact_iff_rTensor_exact, LinearMap.rTensor_exact_iff_lTensor_exact] + +end fixed_universe + +end exact + end FaithfullyFlat end Module From 9248301da2c228348b194c74be94b0e0b0bbbb2c Mon Sep 17 00:00:00 2001 From: Dagur Asgeirsson Date: Wed, 23 Oct 2024 13:28:04 +0000 Subject: [PATCH 333/505] chore(Profinite/Extend): golf and remove unnecessary have statements (#18107) --- Mathlib/Topology/Category/Profinite/Extend.lean | 10 ++-------- 1 file changed, 2 insertions(+), 8 deletions(-) diff --git a/Mathlib/Topology/Category/Profinite/Extend.lean b/Mathlib/Topology/Category/Profinite/Extend.lean index 2c00a5e5808e0..bbdcf1c6b88fc 100644 --- a/Mathlib/Topology/Category/Profinite/Extend.lean +++ b/Mathlib/Topology/Category/Profinite/Extend.lean @@ -126,12 +126,7 @@ def cone (S : Profinite) : pt := G.obj S π := { app := fun i ↦ G.map i.hom - naturality := fun _ _ f ↦ (by - have := f.w - simp only [const_obj_obj, StructuredArrow.left_eq_id, const_obj_map, Category.id_comp, - StructuredArrow.w] at this - simp only [const_obj_obj, comp_obj, StructuredArrow.proj_obj, const_obj_map, Category.id_comp, - Functor.comp_map, StructuredArrow.proj_map, ← map_comp, StructuredArrow.w]) } + naturality := fun _ _ f ↦ (by simp [← map_comp]) } example : G.mapCone c = (cone G c.pt).whisker (functor c) := rfl @@ -166,8 +161,7 @@ def cocone (S : Profinite) : have := f.w simp only [op_obj, const_obj_obj, op_map, CostructuredArrow.right_eq_id, const_obj_map, Category.comp_id] at this - simp only [comp_obj, CostructuredArrow.proj_obj, op_obj, const_obj_obj, Functor.comp_map, - CostructuredArrow.proj_map, op_map, ← map_comp, this, const_obj_map, Category.comp_id]) } + simp [← map_comp, this]) } example : G.mapCocone c.op = (cocone G c.pt).whisker (functorOp c) := rfl From 52eed1022347515addc361c9defb375fdf28c84f Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Wed, 23 Oct 2024 14:34:30 +0000 Subject: [PATCH 334/505] chore(MeasureTheory): delete "`simp` can prove this" porting notes (#18123) This PR checks all porting notes of the form "`simp` can prove this" and if this is indeed the case (and the proof appears to make sense), removes the associated porting note. --- Mathlib/MeasureTheory/Integral/Average.lean | 1 - Mathlib/MeasureTheory/Integral/Lebesgue.lean | 1 - Mathlib/MeasureTheory/Measure/AEMeasurable.lean | 2 +- Mathlib/MeasureTheory/Measure/Count.lean | 2 -- Mathlib/MeasureTheory/Measure/Lebesgue/Basic.lean | 2 -- Mathlib/MeasureTheory/Measure/NullMeasurable.lean | 2 -- Mathlib/MeasureTheory/Measure/Restrict.lean | 1 - 7 files changed, 1 insertion(+), 10 deletions(-) diff --git a/Mathlib/MeasureTheory/Integral/Average.lean b/Mathlib/MeasureTheory/Integral/Average.lean index 86ca2d01fcaf8..7f6337976b25d 100644 --- a/Mathlib/MeasureTheory/Integral/Average.lean +++ b/Mathlib/MeasureTheory/Integral/Average.lean @@ -396,7 +396,6 @@ theorem setAverage_const {s : Set α} (hs₀ : μ s ≠ 0) (hs : μ s ≠ ∞) ( ⨍ _ in s, c ∂μ = c := have := NeZero.mk hs₀; have := Fact.mk hs.lt_top; average_const _ _ --- Porting note (#10618): was `@[simp]` but `simp` can prove it theorem integral_average (μ : Measure α) [IsFiniteMeasure μ] (f : α → E) : ∫ _, ⨍ a, f a ∂μ ∂μ = ∫ x, f x ∂μ := by simp diff --git a/Mathlib/MeasureTheory/Integral/Lebesgue.lean b/Mathlib/MeasureTheory/Integral/Lebesgue.lean index b30251af343f7..2108802bc1e8e 100644 --- a/Mathlib/MeasureTheory/Integral/Lebesgue.lean +++ b/Mathlib/MeasureTheory/Integral/Lebesgue.lean @@ -129,7 +129,6 @@ theorem lintegral_zero : ∫⁻ _ : α, 0 ∂μ = 0 := by simp theorem lintegral_zero_fun : lintegral μ (0 : α → ℝ≥0∞) = 0 := lintegral_zero --- @[simp] -- Porting note (#10618): simp can prove this theorem lintegral_one : ∫⁻ _, (1 : ℝ≥0∞) ∂μ = μ univ := by rw [lintegral_const, one_mul] theorem setLIntegral_const (s : Set α) (c : ℝ≥0∞) : ∫⁻ _ in s, c ∂μ = c * μ s := by diff --git a/Mathlib/MeasureTheory/Measure/AEMeasurable.lean b/Mathlib/MeasureTheory/Measure/AEMeasurable.lean index e1a65cd8f49a1..6f771b3602413 100644 --- a/Mathlib/MeasureTheory/Measure/AEMeasurable.lean +++ b/Mathlib/MeasureTheory/Measure/AEMeasurable.lean @@ -248,7 +248,7 @@ theorem aemeasurable_restrict_iff_comap_subtype {s : Set α} (hs : MeasurableSet {f : α → β} : AEMeasurable f (μ.restrict s) ↔ AEMeasurable (f ∘ (↑) : s → β) (comap (↑) μ) := by rw [← map_comap_subtype_coe hs, (MeasurableEmbedding.subtype_coe hs).aemeasurable_map_iff] -@[to_additive] -- @[to_additive (attr := simp)] -- Porting note (#10618): simp can prove this +@[to_additive] theorem aemeasurable_one [One β] : AEMeasurable (fun _ : α => (1 : β)) μ := measurable_one.aemeasurable diff --git a/Mathlib/MeasureTheory/Measure/Count.lean b/Mathlib/MeasureTheory/Measure/Count.lean index 121403addfb48..b91e37e621e4a 100644 --- a/Mathlib/MeasureTheory/Measure/Count.lean +++ b/Mathlib/MeasureTheory/Measure/Count.lean @@ -37,7 +37,6 @@ theorem le_count_apply : ∑' _ : s, (1 : ℝ≥0∞) ≤ count s := theorem count_apply (hs : MeasurableSet s) : count s = ∑' _ : s, 1 := by simp only [count, sum_apply, hs, dirac_apply', ← tsum_subtype s (1 : α → ℝ≥0∞), Pi.one_apply] --- @[simp] -- Porting note (#10618): simp can prove this theorem count_empty : count (∅ : Set α) = 0 := by rw [count_apply MeasurableSet.empty, tsum_empty] @[simp] @@ -133,7 +132,6 @@ theorem count_singleton' {a : α} (ha : MeasurableSet ({a} : Set α)) : count ({ simp [@toFinset_card _ _ (Set.finite_singleton a).fintype, @Fintype.card_unique _ _ (Set.finite_singleton a).fintype] --- @[simp] -- Porting note (#10618): simp can prove this theorem count_singleton [MeasurableSingletonClass α] (a : α) : count ({a} : Set α) = 1 := count_singleton' (measurableSet_singleton a) diff --git a/Mathlib/MeasureTheory/Measure/Lebesgue/Basic.lean b/Mathlib/MeasureTheory/Measure/Lebesgue/Basic.lean index afb91c570543c..36b280ec89fdf 100644 --- a/Mathlib/MeasureTheory/Measure/Lebesgue/Basic.lean +++ b/Mathlib/MeasureTheory/Measure/Lebesgue/Basic.lean @@ -83,10 +83,8 @@ theorem volume_Ioo {a b : ℝ} : volume (Ioo a b) = ofReal (b - a) := by simp [v @[simp] theorem volume_Ioc {a b : ℝ} : volume (Ioc a b) = ofReal (b - a) := by simp [volume_val] --- @[simp] -- Porting note (#10618): simp can prove this theorem volume_singleton {a : ℝ} : volume ({a} : Set ℝ) = 0 := by simp [volume_val] --- @[simp] -- Porting note (#10618): simp can prove this, after mathlib4#4628 theorem volume_univ : volume (univ : Set ℝ) = ∞ := ENNReal.eq_top_of_forall_nnreal_le fun r => calc diff --git a/Mathlib/MeasureTheory/Measure/NullMeasurable.lean b/Mathlib/MeasureTheory/Measure/NullMeasurable.lean index 6e3eb67afb019..45d1deff59246 100644 --- a/Mathlib/MeasureTheory/Measure/NullMeasurable.lean +++ b/Mathlib/MeasureTheory/Measure/NullMeasurable.lean @@ -92,11 +92,9 @@ def NullMeasurableSet [MeasurableSpace α] (s : Set α) theorem _root_.MeasurableSet.nullMeasurableSet (h : MeasurableSet s) : NullMeasurableSet s μ := h.eventuallyMeasurableSet --- @[simp] -- Porting note (#10618): simp can prove this theorem nullMeasurableSet_empty : NullMeasurableSet ∅ μ := MeasurableSet.empty --- @[simp] -- Porting note (#10618): simp can prove this theorem nullMeasurableSet_univ : NullMeasurableSet univ μ := MeasurableSet.univ diff --git a/Mathlib/MeasureTheory/Measure/Restrict.lean b/Mathlib/MeasureTheory/Measure/Restrict.lean index 5e3e7702c9a7e..70efb6ae3f55d 100644 --- a/Mathlib/MeasureTheory/Measure/Restrict.lean +++ b/Mathlib/MeasureTheory/Measure/Restrict.lean @@ -651,7 +651,6 @@ theorem ae_restrict_eq (hs : MeasurableSet s) : ae (μ.restrict s) = ae μ ⊓ Classical.not_imp, fun a => and_comm (a := a ∈ s) (b := ¬a ∈ t)] rfl --- @[simp] -- Porting note (#10618): simp can prove this theorem ae_restrict_eq_bot {s} : ae (μ.restrict s) = ⊥ ↔ μ s = 0 := ae_eq_bot.trans restrict_eq_zero From 918deb4c8b1f5dad5663d57c158cb37981a446b3 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Wed, 23 Oct 2024 15:11:09 +0000 Subject: [PATCH 335/505] chore(NumberTheory): delete "`simp` can prove this" porting notes (#18124) This PR checks all porting notes of the form "`simp` can prove this" and if this is indeed the case (and the proof appears to make sense), removes the associated porting note. --- .../NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean | 1 - Mathlib/NumberTheory/MulChar/Basic.lean | 1 - Mathlib/NumberTheory/Padics/PadicNorm.lean | 1 - Mathlib/NumberTheory/PellMatiyasevic.lean | 2 -- Mathlib/NumberTheory/Zsqrtd/GaussianInt.lean | 6 ------ 5 files changed, 11 deletions(-) diff --git a/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean b/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean index a46bc6602b64d..b25bd5ca4f8ea 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean @@ -129,7 +129,6 @@ variable {F} theorem quadraticChar_eq_zero_iff {a : F} : quadraticChar F a = 0 ↔ a = 0 := quadraticCharFun_eq_zero_iff --- @[simp] -- Porting note (#10618): simp can prove this theorem quadraticChar_zero : quadraticChar F 0 = 0 := by simp only [quadraticChar_apply, quadraticCharFun_zero] diff --git a/Mathlib/NumberTheory/MulChar/Basic.lean b/Mathlib/NumberTheory/MulChar/Basic.lean index 827d48e679919..925ac23039f92 100644 --- a/Mathlib/NumberTheory/MulChar/Basic.lean +++ b/Mathlib/NumberTheory/MulChar/Basic.lean @@ -311,7 +311,6 @@ theorem inv_apply' {R : Type*} [Field R] (χ : MulChar R R') (a : R) : χ⁻¹ a (inv_apply χ a).trans <| congr_arg _ (Ring.inverse_eq_inv a) /-- The product of a character with its inverse is the trivial character. -/ --- Porting note (#10618): @[simp] can prove this (later) theorem inv_mul (χ : MulChar R R') : χ⁻¹ * χ = 1 := by ext x rw [coeToFun_mul, Pi.mul_apply, inv_apply_eq_inv] diff --git a/Mathlib/NumberTheory/Padics/PadicNorm.lean b/Mathlib/NumberTheory/Padics/PadicNorm.lean index 06bf72812b275..a50570c41e9f1 100644 --- a/Mathlib/NumberTheory/Padics/PadicNorm.lean +++ b/Mathlib/NumberTheory/Padics/PadicNorm.lean @@ -64,7 +64,6 @@ protected theorem nonneg (q : ℚ) : 0 ≤ padicNorm p q := protected theorem zero : padicNorm p 0 = 0 := by simp [padicNorm] /-- The `p`-adic norm of `1` is `1`. -/ --- @[simp] -- Porting note (#10618): simp can prove this protected theorem one : padicNorm p 1 = 1 := by simp [padicNorm] /-- The `p`-adic norm of `p` is `p⁻¹` if `p > 1`. diff --git a/Mathlib/NumberTheory/PellMatiyasevic.lean b/Mathlib/NumberTheory/PellMatiyasevic.lean index d2cef3a69a1f7..8d9f990a8b7df 100644 --- a/Mathlib/NumberTheory/PellMatiyasevic.lean +++ b/Mathlib/NumberTheory/PellMatiyasevic.lean @@ -129,10 +129,8 @@ theorem xn_succ (n : ℕ) : xn a1 (n + 1) = xn a1 n * a + d a1 * yn a1 n := theorem yn_succ (n : ℕ) : yn a1 (n + 1) = xn a1 n + yn a1 n * a := rfl ---@[simp] Porting note (#10618): `simp` can prove it theorem xn_one : xn a1 1 = a := by simp ---@[simp] Porting note (#10618): `simp` can prove it theorem yn_one : yn a1 1 = 1 := by simp /-- The Pell `x` sequence, considered as an integer sequence. -/ diff --git a/Mathlib/NumberTheory/Zsqrtd/GaussianInt.lean b/Mathlib/NumberTheory/Zsqrtd/GaussianInt.lean index c0a2b23beb793..a2da9ffabe073 100644 --- a/Mathlib/NumberTheory/Zsqrtd/GaussianInt.lean +++ b/Mathlib/NumberTheory/Zsqrtd/GaussianInt.lean @@ -89,27 +89,21 @@ theorem toComplex_re (x y : ℤ) : ((⟨x, y⟩ : ℤ[i]) : ℂ).re = x := by si @[simp] theorem toComplex_im (x y : ℤ) : ((⟨x, y⟩ : ℤ[i]) : ℂ).im = y := by simp [toComplex_def] --- Porting note (#10618): @[simp] can prove this theorem toComplex_add (x y : ℤ[i]) : ((x + y : ℤ[i]) : ℂ) = x + y := toComplex.map_add _ _ --- Porting note (#10618): @[simp] can prove this theorem toComplex_mul (x y : ℤ[i]) : ((x * y : ℤ[i]) : ℂ) = x * y := toComplex.map_mul _ _ --- Porting note (#10618): @[simp] can prove this theorem toComplex_one : ((1 : ℤ[i]) : ℂ) = 1 := toComplex.map_one --- Porting note (#10618): @[simp] can prove this theorem toComplex_zero : ((0 : ℤ[i]) : ℂ) = 0 := toComplex.map_zero --- Porting note (#10618): @[simp] can prove this theorem toComplex_neg (x : ℤ[i]) : ((-x : ℤ[i]) : ℂ) = -x := toComplex.map_neg _ --- Porting note (#10618): @[simp] can prove this theorem toComplex_sub (x y : ℤ[i]) : ((x - y : ℤ[i]) : ℂ) = x - y := toComplex.map_sub _ _ From 3b0cdabbff6fe377625a30d838785614fbb71ad9 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Wed, 23 Oct 2024 15:18:39 +0000 Subject: [PATCH 336/505] chore(RingTheory): delete "`simp` can prove this" porting notes (#18119) This PR checks all porting notes of the form "`simp` can prove this" and if this is indeed the case (and the proof appears to make sense), removes the associated porting note. --- Mathlib/RingTheory/AlgebraicIndependent.lean | 2 -- Mathlib/RingTheory/DiscreteValuationRing/Basic.lean | 3 --- .../RingTheory/GradedAlgebra/HomogeneousIdeal.lean | 2 -- Mathlib/RingTheory/HahnSeries/Basic.lean | 1 - Mathlib/RingTheory/HahnSeries/Multiplication.lean | 2 -- Mathlib/RingTheory/Ideal/Operations.lean | 2 -- Mathlib/RingTheory/IsAdjoinRoot.lean | 1 - Mathlib/RingTheory/LaurentSeries.lean | 11 ++++------- Mathlib/RingTheory/Localization/Basic.lean | 4 ---- Mathlib/RingTheory/Localization/Defs.lean | 1 - Mathlib/RingTheory/Localization/Submodule.lean | 1 - Mathlib/RingTheory/Polynomial/Cyclotomic/Eval.lean | 2 -- Mathlib/RingTheory/Polynomial/Hermite/Basic.lean | 2 -- Mathlib/RingTheory/PowerBasis.lean | 1 - Mathlib/RingTheory/RingInvo.lean | 2 -- Mathlib/RingTheory/Valuation/Basic.lean | 5 ----- 16 files changed, 4 insertions(+), 38 deletions(-) diff --git a/Mathlib/RingTheory/AlgebraicIndependent.lean b/Mathlib/RingTheory/AlgebraicIndependent.lean index 43e4a7ea67e93..a161c703a9c61 100644 --- a/Mathlib/RingTheory/AlgebraicIndependent.lean +++ b/Mathlib/RingTheory/AlgebraicIndependent.lean @@ -371,14 +371,12 @@ theorem AlgebraicIndependent.mvPolynomialOptionEquivPolynomialAdjoin_C IsScalarTower.algebraMap_apply R (MvPolynomial ι R), ← Polynomial.C_eq_algebraMap, Polynomial.map_C, RingHom.coe_coe, AlgEquiv.commutes] ---@[simp] Porting note (#10618): simp can prove it theorem AlgebraicIndependent.mvPolynomialOptionEquivPolynomialAdjoin_X_none (hx : AlgebraicIndependent R x) : hx.mvPolynomialOptionEquivPolynomialAdjoin (X none) = Polynomial.X := by rw [AlgebraicIndependent.mvPolynomialOptionEquivPolynomialAdjoin_apply, aeval_X, Option.elim, Polynomial.map_X] ---@[simp] Porting note (#10618): simp can prove it theorem AlgebraicIndependent.mvPolynomialOptionEquivPolynomialAdjoin_X_some (hx : AlgebraicIndependent R x) (i) : hx.mvPolynomialOptionEquivPolynomialAdjoin (X (some i)) = diff --git a/Mathlib/RingTheory/DiscreteValuationRing/Basic.lean b/Mathlib/RingTheory/DiscreteValuationRing/Basic.lean index 74b2f5eeab2fa..8257492736d5b 100644 --- a/Mathlib/RingTheory/DiscreteValuationRing/Basic.lean +++ b/Mathlib/RingTheory/DiscreteValuationRing/Basic.lean @@ -384,11 +384,9 @@ theorem addVal_def' (u : Rˣ) {ϖ : R} (hϖ : Irreducible ϖ) (n : ℕ) : addVal R ((u : R) * ϖ ^ n) = n := addVal_def _ u hϖ n rfl ---@[simp] Porting note (#10618): simp can prove it theorem addVal_zero : addVal R 0 = ⊤ := (addVal R).map_zero ---@[simp] Porting note (#10618): simp can prove it theorem addVal_one : addVal R 1 = 0 := (addVal R).map_one @@ -397,7 +395,6 @@ theorem addVal_uniformizer {ϖ : R} (hϖ : Irreducible ϖ) : addVal R ϖ = 1 := simpa only [one_mul, eq_self_iff_true, Units.val_one, pow_one, forall_true_left, Nat.cast_one] using addVal_def ϖ 1 hϖ 1 ---@[simp] Porting note (#10618): simp can prove it theorem addVal_mul {a b : R} : addVal R (a * b) = addVal R a + addVal R b := (addVal R).map_mul _ _ diff --git a/Mathlib/RingTheory/GradedAlgebra/HomogeneousIdeal.lean b/Mathlib/RingTheory/GradedAlgebra/HomogeneousIdeal.lean index bd67fd0c74783..a87561a7ee276 100644 --- a/Mathlib/RingTheory/GradedAlgebra/HomogeneousIdeal.lean +++ b/Mathlib/RingTheory/GradedAlgebra/HomogeneousIdeal.lean @@ -359,12 +359,10 @@ theorem toIdeal_iInf {κ : Sort*} (s : κ → HomogeneousIdeal 𝒜) : (⨅ i, s i).toIdeal = ⨅ i, (s i).toIdeal := by rw [iInf, toIdeal_sInf, iInf_range] --- @[simp] -- Porting note (#10618): simp can prove this theorem toIdeal_iSup₂ {κ : Sort*} {κ' : κ → Sort*} (s : ∀ i, κ' i → HomogeneousIdeal 𝒜) : (⨆ (i) (j), s i j).toIdeal = ⨆ (i) (j), (s i j).toIdeal := by simp_rw [toIdeal_iSup] --- @[simp] -- Porting note (#10618): simp can prove this theorem toIdeal_iInf₂ {κ : Sort*} {κ' : κ → Sort*} (s : ∀ i, κ' i → HomogeneousIdeal 𝒜) : (⨅ (i) (j), s i j).toIdeal = ⨅ (i) (j), (s i j).toIdeal := by simp_rw [toIdeal_iInf] diff --git a/Mathlib/RingTheory/HahnSeries/Basic.lean b/Mathlib/RingTheory/HahnSeries/Basic.lean index 7c4261e42f42c..74452d2da238a 100644 --- a/Mathlib/RingTheory/HahnSeries/Basic.lean +++ b/Mathlib/RingTheory/HahnSeries/Basic.lean @@ -194,7 +194,6 @@ theorem support_single_subset : support (single a r) ⊆ {a} := by theorem eq_of_mem_support_single {b : Γ} (h : b ∈ support (single a r)) : b = a := support_single_subset h ---@[simp] Porting note (#10618): simp can prove it theorem single_eq_zero : single a (0 : R) = 0 := (single a).map_zero diff --git a/Mathlib/RingTheory/HahnSeries/Multiplication.lean b/Mathlib/RingTheory/HahnSeries/Multiplication.lean index 8f62c2deef2dc..7719c9ebf8a30 100644 --- a/Mathlib/RingTheory/HahnSeries/Multiplication.lean +++ b/Mathlib/RingTheory/HahnSeries/Multiplication.lean @@ -653,11 +653,9 @@ def C : R →+* HahnSeries Γ R where by_cases h : a = 0 <;> simp [h] map_mul' x y := by rw [single_mul_single, zero_add] ---@[simp] Porting note (#10618): simp can prove it theorem C_zero : C (0 : R) = (0 : HahnSeries Γ R) := C.map_zero ---@[simp] Porting note (#10618): simp can prove it theorem C_one : C (1 : R) = (1 : HahnSeries Γ R) := C.map_one diff --git a/Mathlib/RingTheory/Ideal/Operations.lean b/Mathlib/RingTheory/Ideal/Operations.lean index fdc097725c421..8a168c723e1d3 100644 --- a/Mathlib/RingTheory/Ideal/Operations.lean +++ b/Mathlib/RingTheory/Ideal/Operations.lean @@ -554,10 +554,8 @@ theorem pow_sup_pow_eq_top {m n : ℕ} (h : I ⊔ J = ⊤) : I ^ m ⊔ J ^ n = variable (I) --- @[simp] -- Porting note (#10618): simp can prove this theorem mul_bot : I * ⊥ = ⊥ := by simp --- @[simp] -- Porting note (#10618): simp can prove thisrove this theorem bot_mul : ⊥ * I = ⊥ := by simp @[simp] diff --git a/Mathlib/RingTheory/IsAdjoinRoot.lean b/Mathlib/RingTheory/IsAdjoinRoot.lean index 1b153b9e33498..32da2faaaf6e3 100644 --- a/Mathlib/RingTheory/IsAdjoinRoot.lean +++ b/Mathlib/RingTheory/IsAdjoinRoot.lean @@ -141,7 +141,6 @@ theorem aeval_eq (h : IsAdjoinRoot S f) (p : R[X]) : aeval h.root p = h.map p := rw [map_mul, aeval_C, map_pow, aeval_X, RingHom.map_mul, ← h.algebraMap_apply, RingHom.map_pow, map_X] --- @[simp] -- Porting note (#10618): simp can prove this theorem aeval_root (h : IsAdjoinRoot S f) : aeval h.root f = 0 := by rw [aeval_eq, map_self] /-- Choose an arbitrary representative so that `h.map (h.repr x) = x`. diff --git a/Mathlib/RingTheory/LaurentSeries.lean b/Mathlib/RingTheory/LaurentSeries.lean index c588752360d21..14f5ca19d8a36 100644 --- a/Mathlib/RingTheory/LaurentSeries.lean +++ b/Mathlib/RingTheory/LaurentSeries.lean @@ -279,15 +279,15 @@ open LaurentSeries variable {R' : Type*} [Semiring R] [Ring R'] (f g : R⟦X⟧) (f' g' : R'⟦X⟧) -@[norm_cast] -- Porting note (#10618): simp can prove this +@[norm_cast] theorem coe_zero : ((0 : R⟦X⟧) : R⸨X⸩) = 0 := (ofPowerSeries ℤ R).map_zero -@[norm_cast] -- Porting note (#10618): simp can prove this +@[norm_cast] theorem coe_one : ((1 : R⟦X⟧) : R⸨X⸩) = 1 := (ofPowerSeries ℤ R).map_one -@[norm_cast] -- Porting note (#10618): simp can prove this +@[norm_cast] theorem coe_add : ((f + g : R⟦X⟧) : R⸨X⸩) = f + g := (ofPowerSeries ℤ R).map_add _ _ @@ -299,7 +299,7 @@ theorem coe_sub : ((f' - g' : R'⟦X⟧) : R'⸨X⸩) = f' - g' := theorem coe_neg : ((-f' : R'⟦X⟧) : R'⸨X⸩) = -f' := (ofPowerSeries ℤ R').map_neg _ -@[norm_cast] -- Porting note (#10618): simp can prove this +@[norm_cast] theorem coe_mul : ((f * g : R⟦X⟧) : R⸨X⸩) = f * g := (ofPowerSeries ℤ R).map_mul _ _ @@ -314,12 +314,9 @@ theorem coeff_coe (i : ℤ) : Ne, toPowerSeries_symm_apply_coeff, mem_support, imp_true_iff, not_false_iff, reduceCtorEq] --- Porting note (#10618): simp can prove this --- Porting note: removed norm_cast attribute theorem coe_C (r : R) : ((C R r : R⟦X⟧) : R⸨X⸩) = HahnSeries.C r := ofPowerSeries_C _ --- @[simp] -- Porting note (#10618): simp can prove this theorem coe_X : ((X : R⟦X⟧) : R⸨X⸩) = single 1 1 := ofPowerSeries_X diff --git a/Mathlib/RingTheory/Localization/Basic.lean b/Mathlib/RingTheory/Localization/Basic.lean index 9778d31970f78..0a056d7ae3e6f 100644 --- a/Mathlib/RingTheory/Localization/Basic.lean +++ b/Mathlib/RingTheory/Localization/Basic.lean @@ -118,11 +118,9 @@ noncomputable def algEquiv : S ≃ₐ[R] Q := end --- Porting note (#10618): removed `simp`, `simp` can prove it theorem algEquiv_mk' (x : R) (y : M) : algEquiv M S Q (mk' S x y) = mk' Q x y := by simp --- Porting note (#10618): removed `simp`, `simp` can prove it theorem algEquiv_symm_mk' (x : R) (y : M) : (algEquiv M S Q).symm (mk' Q x y) = mk' S x y := by simp variable (M) in @@ -265,11 +263,9 @@ noncomputable def _root_.IsLocalization.unique (R Rₘ) [CommSemiring R] [CommSe end --- Porting note (#10618): removed `simp`, `simp` can prove it nonrec theorem algEquiv_mk' (x : R) (y : M) : algEquiv M S (mk' (Localization M) x y) = mk' S x y := algEquiv_mk' _ _ --- Porting note (#10618): removed `simp`, `simp` can prove it nonrec theorem algEquiv_symm_mk' (x : R) (y : M) : (algEquiv M S).symm (mk' S x y) = mk' (Localization M) x y := algEquiv_symm_mk' _ _ diff --git a/Mathlib/RingTheory/Localization/Defs.lean b/Mathlib/RingTheory/Localization/Defs.lean index 4a4ae21fb0303..3d6a50bf16c29 100644 --- a/Mathlib/RingTheory/Localization/Defs.lean +++ b/Mathlib/RingTheory/Localization/Defs.lean @@ -639,7 +639,6 @@ theorem ringEquivOfRingEquiv_eq_map {j : R ≃+* P} (H : M.map j.toMonoidHom = T map Q (j : R →+* P) (M.le_comap_of_map_le (le_of_eq H)) := rfl --- Porting note (#10618): removed `simp`, `simp` can prove it theorem ringEquivOfRingEquiv_eq {j : R ≃+* P} (H : M.map j.toMonoidHom = T) (x) : ringEquivOfRingEquiv S Q j H ((algebraMap R S) x) = algebraMap P Q (j x) := by simp diff --git a/Mathlib/RingTheory/Localization/Submodule.lean b/Mathlib/RingTheory/Localization/Submodule.lean index 1e3d73e6d705b..1d9412e648572 100644 --- a/Mathlib/RingTheory/Localization/Submodule.lean +++ b/Mathlib/RingTheory/Localization/Submodule.lean @@ -67,7 +67,6 @@ theorem coeSubmodule_span (s : Set R) : rw [IsLocalization.coeSubmodule, Ideal.span, Submodule.map_span] rfl --- @[simp] -- Porting note (#10618): simp can prove this theorem coeSubmodule_span_singleton (x : R) : coeSubmodule S (Ideal.span {x}) = Submodule.span R {(algebraMap R S) x} := by rw [coeSubmodule_span, Set.image_singleton] diff --git a/Mathlib/RingTheory/Polynomial/Cyclotomic/Eval.lean b/Mathlib/RingTheory/Polynomial/Cyclotomic/Eval.lean index a0b678d3d5f3f..4729d0aab257a 100644 --- a/Mathlib/RingTheory/Polynomial/Cyclotomic/Eval.lean +++ b/Mathlib/RingTheory/Polynomial/Cyclotomic/Eval.lean @@ -29,7 +29,6 @@ theorem eval_one_cyclotomic_prime {R : Type*} [CommRing R] {p : ℕ} [hn : Fact simp only [cyclotomic_prime, eval_X, one_pow, Finset.sum_const, eval_pow, eval_finset_sum, Finset.card_range, smul_one_eq_cast] --- @[simp] -- Porting note (#10618): simp already proves this theorem eval₂_one_cyclotomic_prime {R S : Type*} [CommRing R] [Semiring S] (f : R →+* S) {p : ℕ} [Fact p.Prime] : eval₂ f 1 (cyclotomic p R) = p := by simp @@ -39,7 +38,6 @@ theorem eval_one_cyclotomic_prime_pow {R : Type*} [CommRing R] {p : ℕ} (k : simp only [cyclotomic_prime_pow_eq_geom_sum hn.out, eval_X, one_pow, Finset.sum_const, eval_pow, eval_finset_sum, Finset.card_range, smul_one_eq_cast] --- @[simp] -- Porting note (#10618): simp already proves this theorem eval₂_one_cyclotomic_prime_pow {R S : Type*} [CommRing R] [Semiring S] (f : R →+* S) {p : ℕ} (k : ℕ) [Fact p.Prime] : eval₂ f 1 (cyclotomic (p ^ (k + 1)) R) = p := by simp diff --git a/Mathlib/RingTheory/Polynomial/Hermite/Basic.lean b/Mathlib/RingTheory/Polynomial/Hermite/Basic.lean index 6f7d2d23edce3..fa81775947129 100644 --- a/Mathlib/RingTheory/Polynomial/Hermite/Basic.lean +++ b/Mathlib/RingTheory/Polynomial/Hermite/Basic.lean @@ -60,8 +60,6 @@ theorem hermite_eq_iterate (n : ℕ) : hermite n = (fun p => X * p - derivative theorem hermite_zero : hermite 0 = C 1 := rfl --- Porting note (#10618): There was initially @[simp] on this line but it was removed --- because simp can prove this theorem theorem hermite_one : hermite 1 = X := by rw [hermite_succ, hermite_zero] simp only [map_one, mul_one, derivative_one, sub_zero] diff --git a/Mathlib/RingTheory/PowerBasis.lean b/Mathlib/RingTheory/PowerBasis.lean index 3800d1cc29033..f410b20ade48b 100644 --- a/Mathlib/RingTheory/PowerBasis.lean +++ b/Mathlib/RingTheory/PowerBasis.lean @@ -442,7 +442,6 @@ noncomputable def map (pb : PowerBasis R S) (e : S ≃ₐ[R] S') : PowerBasis R variable [Algebra A S] [Algebra A S'] --- @[simp] -- Porting note (#10618): simp can prove this theorem minpolyGen_map (pb : PowerBasis A S) (e : S ≃ₐ[A] S') : (pb.map e).minpolyGen = pb.minpolyGen := by dsimp only [minpolyGen, map_dim] diff --git a/Mathlib/RingTheory/RingInvo.lean b/Mathlib/RingTheory/RingInvo.lean index 86cbcf2eb18f4..70f40d0a35201 100644 --- a/Mathlib/RingTheory/RingInvo.lean +++ b/Mathlib/RingTheory/RingInvo.lean @@ -95,8 +95,6 @@ theorem involution (f : RingInvo R) (x : R) : (f (f x).unop).unop = x := theorem coe_ringEquiv (f : RingInvo R) (a : R) : (f : R ≃+* Rᵐᵒᵖ) a = f a := rfl --- porting note (#10618): simp can prove this --- @[simp] theorem map_eq_zero_iff (f : RingInvo R) {x : R} : f x = 0 ↔ x = 0 := f.toRingEquiv.map_eq_zero_iff diff --git a/Mathlib/RingTheory/Valuation/Basic.lean b/Mathlib/RingTheory/Valuation/Basic.lean index eec13a2b17f5f..54e29458ccc66 100644 --- a/Mathlib/RingTheory/Valuation/Basic.lean +++ b/Mathlib/RingTheory/Valuation/Basic.lean @@ -141,15 +141,12 @@ variable (v : Valuation R Γ₀) {x y z : R} @[simp, norm_cast] theorem coe_coe : ⇑(v : R →*₀ Γ₀) = v := rfl --- @[simp] Porting note (#10618): simp can prove this theorem map_zero : v 0 = 0 := v.map_zero' --- @[simp] Porting note (#10618): simp can prove this theorem map_one : v 1 = 1 := v.map_one' --- @[simp] Porting note (#10618): simp can prove this theorem map_mul : ∀ x y, v (x * y) = v x * v y := v.map_mul' @@ -189,7 +186,6 @@ theorem map_sum_lt' {ι : Type*} {s : Finset ι} {f : ι → R} {g : Γ₀} (hg (hf : ∀ i ∈ s, v (f i) < g) : v (∑ i ∈ s, f i) < g := v.map_sum_lt (ne_of_gt hg) hf --- @[simp] Porting note (#10618): simp can prove this theorem map_pow : ∀ (x) (n : ℕ), v (x ^ n) = v x ^ n := v.toMonoidWithZeroHom.toMonoidHom.map_pow @@ -200,7 +196,6 @@ def toPreorder : Preorder R := Preorder.lift v /-- If `v` is a valuation on a division ring then `v(x) = 0` iff `x = 0`. -/ --- @[simp] Porting note (#10618): simp can prove this theorem zero_iff [Nontrivial Γ₀] (v : Valuation K Γ₀) {x : K} : v x = 0 ↔ x = 0 := map_eq_zero v From d122235286a558644e97d58544b34e49463cd1e1 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Wed, 23 Oct 2024 15:18:40 +0000 Subject: [PATCH 337/505] chore(Topology): delete "`simp` can prove this" porting notes (#18120) This PR checks all porting notes of the form "`simp` can prove this" and if this is indeed the case (and the proof appears to make sense), removes the associated porting note. --- Mathlib/Topology/Algebra/Module/Basic.lean | 8 -------- Mathlib/Topology/Algebra/Module/FiniteDimension.lean | 1 - Mathlib/Topology/Basic.lean | 1 - Mathlib/Topology/FiberBundle/Basic.lean | 1 - Mathlib/Topology/Instances/AddCircle.lean | 4 ---- Mathlib/Topology/MetricSpace/HausdorffDistance.lean | 2 -- 6 files changed, 17 deletions(-) diff --git a/Mathlib/Topology/Algebra/Module/Basic.lean b/Mathlib/Topology/Algebra/Module/Basic.lean index 943df043c2c09..c2feb47fccf1a 100644 --- a/Mathlib/Topology/Algebra/Module/Basic.lean +++ b/Mathlib/Topology/Algebra/Module/Basic.lean @@ -378,7 +378,6 @@ instance continuousSemilinearMapClass : map_continuous f := f.2 map_smulₛₗ f := f.toLinearMap.map_smul' --- porting note (#10618): was `simp`, now `simp only` proves it theorem coe_mk (f : M₁ →ₛₗ[σ₁₂] M₂) (h) : (mk f h : M₁ →ₛₗ[σ₁₂] M₂) = f := rfl @@ -441,7 +440,6 @@ protected theorem map_add (f : M₁ →SL[σ₁₂] M₂) (x y : M₁) : f (x + protected theorem map_smulₛₗ (f : M₁ →SL[σ₁₂] M₂) (c : R₁) (x : M₁) : f (c • x) = σ₁₂ c • f x := (toLinearMap _).map_smulₛₗ _ _ --- @[simp] -- Porting note (#10618): simp can prove this protected theorem map_smul [Module R₁ M₂] (f : M₁ →L[R₁] M₂) (c : R₁) (x : M₁) : f (c • x) = c • f x := by simp only [RingHom.id_apply, ContinuousLinearMap.map_smulₛₗ] @@ -1677,11 +1675,9 @@ theorem map_nhds_eq (e : M₁ ≃SL[σ₁₂] M₂) (x : M₁) : map e (𝓝 x) e.toHomeomorph.map_nhds_eq x -- Make some straightforward lemmas available to `simp`. --- @[simp] -- Porting note (#10618): simp can prove this theorem map_zero (e : M₁ ≃SL[σ₁₂] M₂) : e (0 : M₁) = 0 := (e : M₁ →SL[σ₁₂] M₂).map_zero --- @[simp] -- Porting note (#10618): simp can prove this theorem map_add (e : M₁ ≃SL[σ₁₂] M₂) (x y : M₁) : e (x + y) = e x + e y := (e : M₁ →SL[σ₁₂] M₂).map_add x y @@ -1689,11 +1685,9 @@ theorem map_add (e : M₁ ≃SL[σ₁₂] M₂) (x y : M₁) : e (x + y) = e x + theorem map_smulₛₗ (e : M₁ ≃SL[σ₁₂] M₂) (c : R₁) (x : M₁) : e (c • x) = σ₁₂ c • e x := (e : M₁ →SL[σ₁₂] M₂).map_smulₛₗ c x --- @[simp] -- Porting note (#10618): simp can prove this theorem map_smul [Module R₁ M₂] (e : M₁ ≃L[R₁] M₂) (c : R₁) (x : M₁) : e (c • x) = c • e x := (e : M₁ →L[R₁] M₂).map_smul c x --- @[simp] -- Porting note (#10618): simp can prove this theorem map_eq_zero_iff (e : M₁ ≃SL[σ₁₂] M₂) {x : M₁} : e x = 0 ↔ x = 0 := e.toLinearEquiv.map_eq_zero_iff @@ -2051,11 +2045,9 @@ variable {R : Type*} [Ring R] {R₂ : Type*} [Ring R₂] {M : Type*} [Topologica variable {σ₁₂ : R →+* R₂} {σ₂₁ : R₂ →+* R} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] --- @[simp] -- Porting note (#10618): simp can prove this theorem map_sub (e : M ≃SL[σ₁₂] M₂) (x y : M) : e (x - y) = e x - e y := (e : M →SL[σ₁₂] M₂).map_sub x y --- @[simp] -- Porting note (#10618): simp can prove this theorem map_neg (e : M ≃SL[σ₁₂] M₂) (x : M) : e (-x) = -e x := (e : M →SL[σ₁₂] M₂).map_neg x diff --git a/Mathlib/Topology/Algebra/Module/FiniteDimension.lean b/Mathlib/Topology/Algebra/Module/FiniteDimension.lean index ccbd57b6b210b..8f9b4584bce46 100644 --- a/Mathlib/Topology/Algebra/Module/FiniteDimension.lean +++ b/Mathlib/Topology/Algebra/Module/FiniteDimension.lean @@ -374,7 +374,6 @@ theorem toLinearEquiv_toContinuousLinearEquiv (e : E ≃ₗ[𝕜] F) : ext x rfl --- Porting note (#10618): @[simp] can prove this theorem toLinearEquiv_toContinuousLinearEquiv_symm (e : E ≃ₗ[𝕜] F) : e.toContinuousLinearEquiv.symm.toLinearEquiv = e.symm := by ext x diff --git a/Mathlib/Topology/Basic.lean b/Mathlib/Topology/Basic.lean index 1c5fc5ac20a5c..b9f811270fa26 100644 --- a/Mathlib/Topology/Basic.lean +++ b/Mathlib/Topology/Basic.lean @@ -1144,7 +1144,6 @@ theorem dense_compl_singleton (x : X) [NeBot (𝓝[≠] x)] : Dense ({x}ᶜ : Se /-- If `x` is not an isolated point of a topological space, then the closure of `{x}ᶜ` is the whole space. -/ --- Porting note (#10618): was a `@[simp]` lemma but `simp` can prove it theorem closure_compl_singleton (x : X) [NeBot (𝓝[≠] x)] : closure {x}ᶜ = (univ : Set X) := (dense_compl_singleton x).closure_eq diff --git a/Mathlib/Topology/FiberBundle/Basic.lean b/Mathlib/Topology/FiberBundle/Basic.lean index 7fc7b01fd62b1..275daa23c07b7 100644 --- a/Mathlib/Topology/FiberBundle/Basic.lean +++ b/Mathlib/Topology/FiberBundle/Basic.lean @@ -671,7 +671,6 @@ theorem mem_localTrivAt_baseSet (b : B) : b ∈ (Z.localTrivAt b).baseSet := by rw [localTrivAt, ← baseSet_at] exact Z.mem_baseSet_at b --- Porting note (#10618): was @[simp, mfld_simps], now `simp` can prove it theorem mk_mem_localTrivAt_source : (⟨b, a⟩ : Z.TotalSpace) ∈ (Z.localTrivAt b).source := by simp only [mfld_simps] diff --git a/Mathlib/Topology/Instances/AddCircle.lean b/Mathlib/Topology/Instances/AddCircle.lean index 4c5e3613f132b..b762d06ff3e79 100644 --- a/Mathlib/Topology/Instances/AddCircle.lean +++ b/Mathlib/Topology/Instances/AddCircle.lean @@ -152,10 +152,6 @@ theorem coe_eq_zero_of_pos_iff (hp : 0 < p) {x : 𝕜} (hx : 0 < x) : theorem coe_period : (p : AddCircle p) = 0 := (QuotientAddGroup.eq_zero_iff p).2 <| mem_zmultiples p -/- Porting note (#10618): `simp` attribute removed because linter reports: -simp can prove this: - by simp only [@mem_zmultiples, @QuotientAddGroup.mk_add_of_mem] --/ theorem coe_add_period (x : 𝕜) : ((x + p : 𝕜) : AddCircle p) = x := by rw [coe_add, ← eq_sub_iff_add_eq', sub_self, coe_period] diff --git a/Mathlib/Topology/MetricSpace/HausdorffDistance.lean b/Mathlib/Topology/MetricSpace/HausdorffDistance.lean index 06799e111037e..0614f4db7f848 100644 --- a/Mathlib/Topology/MetricSpace/HausdorffDistance.lean +++ b/Mathlib/Topology/MetricSpace/HausdorffDistance.lean @@ -378,7 +378,6 @@ theorem hausdorffEdist_closure₂ : hausdorffEdist s (closure t) = hausdorffEdis simp [@hausdorffEdist_comm _ _ s _] /-- The Hausdorff edistance between sets or their closures is the same. -/ --- @[simp] -- Porting note (#10618): simp can prove this theorem hausdorffEdist_closure : hausdorffEdist (closure s) (closure t) = hausdorffEdist s t := by simp @@ -777,7 +776,6 @@ theorem hausdorffDist_closure₂ : hausdorffDist s (closure t) = hausdorffDist s simp [hausdorffDist] /-- The Hausdorff distances between two sets and their closures coincide. -/ --- @[simp] -- Porting note (#10618): simp can prove this theorem hausdorffDist_closure : hausdorffDist (closure s) (closure t) = hausdorffDist s t := by simp [hausdorffDist] From 9ad8de0a458fe6eb93011f774519e565604a7f2f Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Wed, 23 Oct 2024 15:18:42 +0000 Subject: [PATCH 338/505] chore(LinearAlgebra): delete "`simp` can prove this" porting notes (#18122) This PR checks all porting notes of the form "`simp` can prove this" and if this is indeed the case (and the proof appears to make sense), removes the associated porting note. --- Mathlib/LinearAlgebra/Alternating/Basic.lean | 2 +- Mathlib/LinearAlgebra/Dimension/Constructions.lean | 2 -- Mathlib/LinearAlgebra/ExteriorAlgebra/Basic.lean | 2 -- Mathlib/LinearAlgebra/Finsupp.lean | 1 - Mathlib/LinearAlgebra/Isomorphisms.lean | 1 - Mathlib/LinearAlgebra/Lagrange.lean | 4 ---- Mathlib/LinearAlgebra/Matrix/Charpoly/Univ.lean | 1 - Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean | 1 - Mathlib/LinearAlgebra/Multilinear/Basic.lean | 2 +- Mathlib/LinearAlgebra/Span.lean | 1 - 10 files changed, 2 insertions(+), 15 deletions(-) diff --git a/Mathlib/LinearAlgebra/Alternating/Basic.lean b/Mathlib/LinearAlgebra/Alternating/Basic.lean index 06184ef9cb523..6fc8b20955e39 100644 --- a/Mathlib/LinearAlgebra/Alternating/Basic.lean +++ b/Mathlib/LinearAlgebra/Alternating/Basic.lean @@ -114,7 +114,7 @@ protected theorem congr_arg (f : M [⋀^ι]→ₗ[R] N) {x y : ι → M} (h : x theorem coe_injective : Injective ((↑) : M [⋀^ι]→ₗ[R] N → (ι → M) → N) := DFunLike.coe_injective -@[norm_cast] -- @[simp] -- Porting note (#10618): simp can prove this +@[norm_cast] theorem coe_inj {f g : M [⋀^ι]→ₗ[R] N} : (f : (ι → M) → N) = g ↔ f = g := coe_injective.eq_iff diff --git a/Mathlib/LinearAlgebra/Dimension/Constructions.lean b/Mathlib/LinearAlgebra/Dimension/Constructions.lean index 3c58916c9b5e9..d03d8646c1451 100644 --- a/Mathlib/LinearAlgebra/Dimension/Constructions.lean +++ b/Mathlib/LinearAlgebra/Dimension/Constructions.lean @@ -203,7 +203,6 @@ theorem rank_matrix' (m n : Type v) [Finite m] [Finite n] : /-- If `m` and `n` are `Fintype` that lie in the same universe as `R`, the rank of `m × n` matrices is `# m * # n`. -/ --- @[simp] -- Porting note (#10618): simp can prove this theorem rank_matrix'' (m n : Type u) [Finite m] [Finite n] : Module.rank R (Matrix m n R) = #m * #n := by simp @@ -298,7 +297,6 @@ theorem Module.finrank_fintype_fun_eq_card : finrank R (η → R) = Fintype.card finrank_eq_of_rank_eq rank_fun' /-- The vector space of functions on `Fin n` has finrank equal to `n`. -/ --- @[simp] -- Porting note (#10618): simp already proves this theorem Module.finrank_fin_fun {n : ℕ} : finrank R (Fin n → R) = n := by simp variable {R} diff --git a/Mathlib/LinearAlgebra/ExteriorAlgebra/Basic.lean b/Mathlib/LinearAlgebra/ExteriorAlgebra/Basic.lean index 258bd811ae3b5..75b12c7d91e55 100644 --- a/Mathlib/LinearAlgebra/ExteriorAlgebra/Basic.lean +++ b/Mathlib/LinearAlgebra/ExteriorAlgebra/Basic.lean @@ -82,14 +82,12 @@ end exteriorPower variable {R} /-- As well as being linear, `ι m` squares to zero. -/ --- @[simp] -- Porting note (#10618): simp can prove this theorem ι_sq_zero (m : M) : ι R m * ι R m = 0 := (CliffordAlgebra.ι_sq_scalar _ m).trans <| map_zero _ section variable {A : Type*} [Semiring A] [Algebra R A] --- @[simp] -- Porting note (#10618): simp can prove this theorem comp_ι_sq_zero (g : ExteriorAlgebra R M →ₐ[R] A) (m : M) : g (ι R m) * g (ι R m) = 0 := by rw [← map_mul, ι_sq_zero, map_zero] diff --git a/Mathlib/LinearAlgebra/Finsupp.lean b/Mathlib/LinearAlgebra/Finsupp.lean index 5371c45091225..d8827696368ab 100644 --- a/Mathlib/LinearAlgebra/Finsupp.lean +++ b/Mathlib/LinearAlgebra/Finsupp.lean @@ -889,7 +889,6 @@ theorem domLCongr_symm {α₁ α₂ : Type*} (f : α₁ ≃ α₂) : ((Finsupp.domLCongr f).symm : (_ →₀ M) ≃ₗ[R] _) = Finsupp.domLCongr f.symm := LinearEquiv.ext fun _ => rfl --- @[simp] -- Porting note (#10618): simp can prove this theorem domLCongr_single {α₁ : Type*} {α₂ : Type*} (e : α₁ ≃ α₂) (i : α₁) (m : M) : (Finsupp.domLCongr e : _ ≃ₗ[R] _) (Finsupp.single i m) = Finsupp.single (e i) m := by simp diff --git a/Mathlib/LinearAlgebra/Isomorphisms.lean b/Mathlib/LinearAlgebra/Isomorphisms.lean index 998bbedc53b52..0dc3fa092cd9c 100644 --- a/Mathlib/LinearAlgebra/Isomorphisms.lean +++ b/Mathlib/LinearAlgebra/Isomorphisms.lean @@ -117,7 +117,6 @@ theorem quotientInfEquivSupQuotient_symm_apply_left (p p' : Submodule R M) (x : rw [quotientInfEquivSupQuotient_apply_mk, inclusion_apply] --- @[simp] -- Porting note (#10618): simp can prove this theorem quotientInfEquivSupQuotient_symm_apply_eq_zero_iff {p p' : Submodule R M} {x : ↥(p ⊔ p')} : (quotientInfEquivSupQuotient p p').symm (Submodule.Quotient.mk x) = 0 ↔ (x : M) ∈ p' := (LinearEquiv.symm_apply_eq _).trans <| by simp diff --git a/Mathlib/LinearAlgebra/Lagrange.lean b/Mathlib/LinearAlgebra/Lagrange.lean index b683703a0116d..eebd330370ef9 100644 --- a/Mathlib/LinearAlgebra/Lagrange.lean +++ b/Mathlib/LinearAlgebra/Lagrange.lean @@ -296,12 +296,8 @@ def interpolate (s : Finset ι) (v : ι → F) : (ι → F) →ₗ[F] F[X] where map_smul' c f := by simp_rw [Finset.smul_sum, C_mul', smul_smul, Pi.smul_apply, RingHom.id_apply, smul_eq_mul] --- Porting note (#10618): There was originally '@[simp]' on this line but it was removed because --- 'simp' could prove 'interpolate_empty' theorem interpolate_empty : interpolate ∅ v r = 0 := by rw [interpolate_apply, sum_empty] --- Porting note (#10618): There was originally '@[simp]' on this line but it was removed because --- 'simp' could prove 'interpolate_singleton' theorem interpolate_singleton : interpolate {i} v r = C (r i) := by rw [interpolate_apply, sum_singleton, basis_singleton, mul_one] diff --git a/Mathlib/LinearAlgebra/Matrix/Charpoly/Univ.lean b/Mathlib/LinearAlgebra/Matrix/Charpoly/Univ.lean index 4f8c88911dcd1..e4c7fb458c0ae 100644 --- a/Mathlib/LinearAlgebra/Matrix/Charpoly/Univ.lean +++ b/Mathlib/LinearAlgebra/Matrix/Charpoly/Univ.lean @@ -71,7 +71,6 @@ variable (R) lemma univ_monic : (univ R n).Monic := charpoly_monic (mvPolynomialX n n R) --- Porting note (#10618): no @[simp], since simp can prove this lemma univ_natDegree [Nontrivial R] : (univ R n).natDegree = Fintype.card n := charpoly_natDegree_eq_dim (mvPolynomialX n n R) diff --git a/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean b/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean index 2780bc5cd41ea..24fa035e5eef1 100644 --- a/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean +++ b/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean @@ -75,7 +75,6 @@ theorem det_diagonal {d : n → R} : det (diagonal d) = ∏ i, d i := by · simp · simp --- @[simp] -- Porting note (#10618): simp can prove this theorem det_zero (_ : Nonempty n) : det (0 : Matrix n n R) = 0 := (detRowAlternating : (n → R) [⋀^n]→ₗ[R] R).map_zero diff --git a/Mathlib/LinearAlgebra/Multilinear/Basic.lean b/Mathlib/LinearAlgebra/Multilinear/Basic.lean index 39f1d7e414a2b..9b8313b2b87e7 100644 --- a/Mathlib/LinearAlgebra/Multilinear/Basic.lean +++ b/Mathlib/LinearAlgebra/Multilinear/Basic.lean @@ -127,7 +127,7 @@ nonrec theorem congr_arg (f : MultilinearMap R M₁ M₂) {x y : ∀ i, M₁ i} theorem coe_injective : Injective ((↑) : MultilinearMap R M₁ M₂ → (∀ i, M₁ i) → M₂) := DFunLike.coe_injective -@[norm_cast] -- Porting note (#10618): Removed simp attribute, simp can prove this +@[norm_cast] theorem coe_inj {f g : MultilinearMap R M₁ M₂} : (f : (∀ i, M₁ i) → M₂) = g ↔ f = g := DFunLike.coe_fn_eq diff --git a/Mathlib/LinearAlgebra/Span.lean b/Mathlib/LinearAlgebra/Span.lean index 478735c139125..9fb5fb671feff 100644 --- a/Mathlib/LinearAlgebra/Span.lean +++ b/Mathlib/LinearAlgebra/Span.lean @@ -947,7 +947,6 @@ theorem span_singleton_eq_range (x : M) : (R ∙ x) = range (toSpanSingleton R M refine Iff.trans ?_ LinearMap.mem_range.symm exact mem_span_singleton --- @[simp] -- Porting note (#10618): simp can prove this theorem toSpanSingleton_one (x : M) : toSpanSingleton R M x 1 = x := one_smul _ _ From 0ca090b93ba400da1018a33d67762bd3f9772389 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Wed, 23 Oct 2024 15:46:59 +0000 Subject: [PATCH 339/505] feat(SetTheory/Ordinal/Rank): in a well-order, `IsWellFounded.rank = Ordinal.typein` (#18079) --- Mathlib/SetTheory/Ordinal/Rank.lean | 31 ++++++++++++++++++++++++++--- 1 file changed, 28 insertions(+), 3 deletions(-) diff --git a/Mathlib/SetTheory/Ordinal/Rank.lean b/Mathlib/SetTheory/Ordinal/Rank.lean index 20aa63340bd04..e4b3808a56625 100644 --- a/Mathlib/SetTheory/Ordinal/Rank.lean +++ b/Mathlib/SetTheory/Ordinal/Rank.lean @@ -22,7 +22,7 @@ namespace Acc variable {r : α → α → Prop} -/-- The rank of an element `a` accessible under a relation `r` is defined inductively as the +/-- The rank of an element `a` accessible under a relation `r` is defined recursively as the smallest ordinal greater than the ranks of all elements below it (i.e. elements `b` such that `r b a`). -/ noncomputable def rank (h : Acc r a) : Ordinal.{u} := @@ -39,6 +39,18 @@ theorem rank_lt_of_rel (hb : Acc r b) (h : r a b) : (hb.inv h).rank < hb.rank := rw [hb.rank_eq] exact Ordinal.le_iSup _ (⟨a, h⟩ : {a // r a b}) +theorem mem_range_rank_of_le {o : Ordinal} (ha : Acc r a) (ho : o ≤ ha.rank) : + ∃ (b : α) (hb : Acc r b), hb.rank = o := by + obtain rfl | ho := ho.eq_or_lt + · exact ⟨a, ha, rfl⟩ + · revert ho + refine ha.recOn fun a ha IH ho ↦ ?_ + rw [rank_eq, Ordinal.lt_iSup] at ho + obtain ⟨⟨b, hb⟩, ho⟩ := ho + rw [Order.lt_succ_iff] at ho + obtain rfl | ho := ho.eq_or_lt + exacts [⟨b, ha b hb, rfl⟩, IH _ hb ho] + end Acc /-! ### Rank in a well-founded relation -/ @@ -47,7 +59,7 @@ namespace IsWellFounded variable (r : α → α → Prop) [hwf : IsWellFounded α r] -/-- The rank of an element `a` under a well-founded relation `r` is defined inductively as the +/-- The rank of an element `a` under a well-founded relation `r` is defined recursively as the smallest ordinal greater than the ranks of all elements below it (i.e. elements `b` such that `r b a`). -/ noncomputable def rank (a : α) : Ordinal.{u} := @@ -56,10 +68,15 @@ noncomputable def rank (a : α) : Ordinal.{u} := theorem rank_eq (a : α) : rank r a = ⨆ b : { b // r b a }, Order.succ (rank r b) := (hwf.apply r a).rank_eq -variable {r} in +variable {r : α → α → Prop} [hwf : IsWellFounded α r] + theorem rank_lt_of_rel (h : r a b) : rank r a < rank r b := Acc.rank_lt_of_rel _ h +theorem mem_range_rank_of_le {o : Ordinal} (h : o ≤ rank r a) : o ∈ Set.range (rank r) := by + obtain ⟨b, hb, rfl⟩ := Acc.mem_range_rank_of_le (hwf.apply r a) h + exact ⟨b, rfl⟩ + end IsWellFounded theorem WellFoundedLT.rank_strictMono [Preorder α] [WellFoundedLT α] : @@ -70,6 +87,14 @@ theorem WellFoundedGT.rank_strictAnti [Preorder α] [WellFoundedGT α] : StrictAnti (IsWellFounded.rank (α := α) (· > ·)) := fun _ _ a => IsWellFounded.rank_lt_of_rel a +@[simp] +theorem IsWellFounded.rank_eq_typein (r) [IsWellOrder α r] : rank r = Ordinal.typein r := by + classical + letI := linearOrderOfSTO r + ext a + exact InitialSeg.eq (⟨(OrderEmbedding.ofStrictMono _ WellFoundedLT.rank_strictMono).ltEmbedding, + fun a b h ↦ mem_range_rank_of_le h.le⟩) (Ordinal.typein r) a + namespace WellFounded set_option linter.deprecated false From 9e3c142985604203ca1c976f9960bd5a2a19999e Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Wed, 23 Oct 2024 15:47:00 +0000 Subject: [PATCH 340/505] chore(Algebra/Category): restore lemmas lost from the port (#18112) Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- Mathlib/Algebra/Category/Ring/Basic.lean | 44 ++++++++++++++++-------- 1 file changed, 30 insertions(+), 14 deletions(-) diff --git a/Mathlib/Algebra/Category/Ring/Basic.lean b/Mathlib/Algebra/Category/Ring/Basic.lean index 801518f945a60..4d03754154342 100644 --- a/Mathlib/Algebra/Category/Ring/Basic.lean +++ b/Mathlib/Algebra/Category/Ring/Basic.lean @@ -139,6 +139,11 @@ theorem ofHom_apply {R S : Type u} [Semiring R] [Semiring S] (f : R →+* S) (x ofHom f x = f x := rfl +/-- A variant of `ofHom_apply` that makes `simpNF` happy -/ +@[simp] +theorem ofHom_apply' {R S : Type u} [Semiring R] [Semiring S] (f : R →+* S) (x : R) : + DFunLike.coe (α := R) (β := fun _ ↦ S) (ofHom f) x = f x := rfl + /-- Ring equivalence are isomorphisms in category of semirings -/ @@ -199,10 +204,9 @@ def of (R : Type u) [Ring R] : RingCat := def ofHom {R S : Type u} [Ring R] [Ring S] (f : R →+* S) : of R ⟶ of S := f --- Porting note: I think this is now redundant. --- @[simp] --- theorem ofHom_apply {R S : Type u} [Ring R] [Ring S] (f : R →+* S) (x : R) : ofHom f x = f x := --- rfl +theorem ofHom_apply {R S : Type u} [Ring R] [Ring S] (f : R →+* S) (x : R) : ofHom f x = f x := + rfl + instance : Inhabited RingCat := ⟨of PUnit⟩ @@ -214,6 +218,11 @@ instance (R : RingCat) : Ring R := theorem coe_of (R : Type u) [Ring R] : (RingCat.of R : Type u) = R := rfl +/-- A variant of `ofHom_apply` that makes `simpNF` happy -/ +@[simp] +theorem ofHom_apply' {R S : Type u} [Ring R] [Ring S] (f : R →+* S) (x : R) : + DFunLike.coe (α := R) (β := fun _ ↦ S) (ofHom f) x = f x := rfl + -- Coercing the identity morphism, as a ring homomorphism, gives the identity function. @[simp] theorem coe_ringHom_id {X : RingCat} : @DFunLike.coe (X →+* X) X (fun _ ↦ X) _ (𝟙 X) = id := @@ -324,11 +333,14 @@ lemma RingEquiv_coe_eq {X Y : Type _} [CommSemiring X] [CommSemiring Y] (e : X ConcreteCategory.instFunLike (e : X →+* Y) : X → Y) = ↑e := rfl --- Porting note: I think this is now redundant. --- @[simp] --- theorem ofHom_apply {R S : Type u} [CommSemiring R] [CommSemiring S] (f : R →+* S) (x : R) : --- ofHom f x = f x := --- rfl +theorem ofHom_apply {R S : Type u} [CommSemiring R] [CommSemiring S] (f : R →+* S) (x : R) : + ofHom f x = f x := + rfl + +/-- A variant of `ofHom_apply` that makes `simpNF` happy -/ +@[simp] +theorem ofHom_apply' {R S : Type u} [CommSemiring R] [CommSemiring S] (f : R →+* S) (x : R) : + DFunLike.coe (α := R) (β := fun _ ↦ S) (ofHom f) x = f x := rfl instance : Inhabited CommSemiRingCat := ⟨of PUnit⟩ @@ -472,11 +484,15 @@ lemma RingEquiv_coe_eq {X Y : Type _} [CommRing X] [CommRing Y] (e : X ≃+* Y) ConcreteCategory.instFunLike (e : X →+* Y) : X → Y) = ↑e := rfl --- Porting note: I think this is now redundant. --- @[simp] --- theorem ofHom_apply {R S : Type u} [CommRing R] [CommRing S] (f : R →+* S) (x : R) : --- ofHom f x = f x := --- rfl +theorem ofHom_apply {R S : Type u} [CommRing R] [CommRing S] (f : R →+* S) (x : R) : + ofHom f x = f x := + rfl + +/-- A variant of `ofHom_apply` that makes `simpNF` happy -/ +@[simp] +theorem ofHom_apply' {R S : Type u} [CommRing R] [CommRing S] (f : R →+* S) (x : R) : + DFunLike.coe (α := R) (β := fun _ ↦ S) (ofHom f) x = f x := rfl + instance : Inhabited CommRingCat := ⟨of PUnit⟩ From 9f64b5d99fa1362a1d1f21b87381f94cee642571 Mon Sep 17 00:00:00 2001 From: Robin Carlier Date: Wed, 23 Oct 2024 16:51:33 +0000 Subject: [PATCH 341/505] feat(CategoryTheory/ChosenFiniteProducts): Monoidal functors between categories with chosen products (#17613) Introduce a counterpart of `Functor.toMonoidalFunctorOfHasFiniteProducts` for categories with chosen finite products. Co-authored-by: Robin Carlier <57142648+robin-carlier@users.noreply.github.com> --- .../CategoryTheory/ChosenFiniteProducts.lean | 50 +++++++++++++++++++ 1 file changed, 50 insertions(+) diff --git a/Mathlib/CategoryTheory/ChosenFiniteProducts.lean b/Mathlib/CategoryTheory/ChosenFiniteProducts.lean index 4fda2a453f164..13bd59c64fac4 100644 --- a/Mathlib/CategoryTheory/ChosenFiniteProducts.lean +++ b/Mathlib/CategoryTheory/ChosenFiniteProducts.lean @@ -5,6 +5,7 @@ Authors: Adam Topaz, Robin Carlier -/ import Mathlib.CategoryTheory.Monoidal.OfChosenFiniteProducts.Symmetric import Mathlib.CategoryTheory.Limits.Constructions.FiniteProductsOfBinaryProducts +import Mathlib.CategoryTheory.Limits.Preserves.Shapes.BinaryProducts /-! # Categories with chosen finite products @@ -484,4 +485,53 @@ end ChosenFiniteProductsComparison end ChosenFiniteProducts +open Limits MonoidalCategory ChosenFiniteProducts + +variable {C : Type u} [Category.{v} C] [ChosenFiniteProducts C] + {D : Type u₁} [Category.{v₁} D] [ChosenFiniteProducts D] (F : C ⥤ D) + +/-- Any functor between categories with chosen finite products induces an oplax monoial functor. -/ +@[simps] +def Functor.toOplaxMonoidalFunctorOfChosenFiniteProducts : OplaxMonoidalFunctor C D where + toFunctor := F + η := terminalComparison F + δ X Y := prodComparison F X Y + δ_natural_left {X Y} f X' := by + symm; simpa using ChosenFiniteProducts.prodComparison_natural F f (𝟙 X') + δ_natural_right {X Y} X' g := by + symm; simpa using ChosenFiniteProducts.prodComparison_natural F (𝟙 X') g + associativity X Y Z := by + apply hom_ext + case' h_snd => apply hom_ext + all_goals simp [← Functor.map_comp] + left_unitality X := by + apply hom_ext + · exact toUnit_unique _ _ + · simp only [leftUnitor_inv_snd, Category.assoc, whiskerRight_snd, + ChosenFiniteProducts.prodComparison_snd, ← F.map_comp, F.map_id] + right_unitality X := by + apply hom_ext + · simp only [rightUnitor_inv_fst, Category.assoc, whiskerLeft_fst, + ChosenFiniteProducts.prodComparison_fst, ← F.map_comp, F.map_id] + · exact toUnit_unique _ _ + +variable [PreservesLimit (Functor.empty.{0} C) F] + [PreservesLimitsOfShape (Discrete WalkingPair) F] + +instance : IsIso F.toOplaxMonoidalFunctorOfChosenFiniteProducts.η := + terminalComparison_isIso_of_preservesLimits F + +instance (A B : C) : IsIso (F.toOplaxMonoidalFunctorOfChosenFiniteProducts.δ A B) := + isIso_prodComparison_of_preservesLimit_pair F A B + +/-- If `F` preserves finite products, the oplax monoidal functor +`F.toOplaxMonoidalFunctorOfChosenFiniteProducts` can be promoted to a strong monoidal functor. -/ +@[simps!] +noncomputable def Functor.toMonoidalFunctorOfChosenFiniteProducts : MonoidalFunctor C D := + MonoidalFunctor.fromOplaxMonoidalFunctor + (toOplaxMonoidalFunctorOfChosenFiniteProducts F) + +instance [F.IsEquivalence] : F.toMonoidalFunctorOfChosenFiniteProducts.IsEquivalence := by + assumption + end CategoryTheory From 8a18755e1f1ed3b93176930891acefa2cdd71df1 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Wed, 23 Oct 2024 16:59:52 +0000 Subject: [PATCH 342/505] chore(Order): delete "`simp` can prove this" porting notes (#18125) This PR checks all porting notes of the form "`simp` can prove this" and if this is indeed the case (and the proof appears to make sense), removes the associated porting note. --- Mathlib/Order/Basic.lean | 2 -- Mathlib/Order/Chain.lean | 2 -- Mathlib/Order/Hom/Basic.lean | 3 -- Mathlib/Order/Ideal.lean | 2 -- Mathlib/Order/Interval/Finset/Basic.lean | 32 ------------------- Mathlib/Order/Interval/Finset/Fin.lean | 16 ---------- Mathlib/Order/Interval/Finset/Nat.lean | 12 ------- Mathlib/Order/Interval/Multiset.lean | 7 ---- Mathlib/Order/Interval/Set/Basic.lean | 30 ----------------- .../Order/Interval/Set/UnorderedInterval.lean | 2 -- Mathlib/Order/Irreducible.lean | 4 --- Mathlib/Order/PartialSups.lean | 1 - 12 files changed, 113 deletions(-) diff --git a/Mathlib/Order/Basic.lean b/Mathlib/Order/Basic.lean index f7a234670ae36..1249ba8d9278a 100644 --- a/Mathlib/Order/Basic.lean +++ b/Mathlib/Order/Basic.lean @@ -1283,11 +1283,9 @@ theorem max_eq : max a b = unit := theorem min_eq : min a b = unit := rfl --- Porting note (#10618): simp can prove this @[simp] protected theorem le : a ≤ b := trivial --- Porting note (#10618): simp can prove this @[simp] theorem not_lt : ¬a < b := not_false diff --git a/Mathlib/Order/Chain.lean b/Mathlib/Order/Chain.lean index f8f5ce2631f67..f83b31886b200 100644 --- a/Mathlib/Order/Chain.lean +++ b/Mathlib/Order/Chain.lean @@ -279,8 +279,6 @@ instance : SetLike (Flag α) α where theorem ext : (s : Set α) = t → s = t := SetLike.ext' --- Porting note (#10618): `simp` can now prove this --- @[simp] theorem mem_coe_iff : a ∈ (s : Set α) ↔ a ∈ s := Iff.rfl diff --git a/Mathlib/Order/Hom/Basic.lean b/Mathlib/Order/Hom/Basic.lean index c9ab0123359da..038a9e1b51ebd 100644 --- a/Mathlib/Order/Hom/Basic.lean +++ b/Mathlib/Order/Hom/Basic.lean @@ -755,8 +755,6 @@ protected theorem injective (e : α ≃o β) : Function.Injective e := protected theorem surjective (e : α ≃o β) : Function.Surjective e := e.toEquiv.surjective --- Porting note (#10618): simp can prove this --- @[simp] theorem apply_eq_iff_eq (e : α ≃o β) {x y : α} : e x = e y ↔ x = y := e.toEquiv.apply_eq_iff_eq @@ -920,7 +918,6 @@ section LE variable [LE α] [LE β] ---@[simp] Porting note (#10618): simp can prove it theorem le_iff_le (e : α ≃o β) {x y : α} : e x ≤ e y ↔ x ≤ y := e.map_rel_iff diff --git a/Mathlib/Order/Ideal.lean b/Mathlib/Order/Ideal.lean index 9fe09de57fd26..4dae232039074 100644 --- a/Mathlib/Order/Ideal.lean +++ b/Mathlib/Order/Ideal.lean @@ -134,11 +134,9 @@ theorem mem_compl_of_ge {x y : P} : x ≤ y → x ∈ (I : Set P)ᶜ → y ∈ ( instance instPartialOrderIdeal : PartialOrder (Ideal P) := PartialOrder.lift SetLike.coe SetLike.coe_injective --- @[simp] -- Porting note (#10618): simp can prove this theorem coe_subset_coe : (s : Set P) ⊆ t ↔ s ≤ t := Iff.rfl --- @[simp] -- Porting note (#10618): simp can prove this theorem coe_ssubset_coe : (s : Set P) ⊂ t ↔ s < t := Iff.rfl diff --git a/Mathlib/Order/Interval/Finset/Basic.lean b/Mathlib/Order/Interval/Finset/Basic.lean index 4f194b2424862..e8a5116a9e300 100644 --- a/Mathlib/Order/Interval/Finset/Basic.lean +++ b/Mathlib/Order/Interval/Finset/Basic.lean @@ -120,36 +120,20 @@ theorem Ioc_eq_empty_of_le (h : b ≤ a) : Ioc a b = ∅ := theorem Ioo_eq_empty_of_le (h : b ≤ a) : Ioo a b = ∅ := Ioo_eq_empty h.not_lt --- porting note (#10618): simp can prove this --- @[simp] theorem left_mem_Icc : a ∈ Icc a b ↔ a ≤ b := by simp only [mem_Icc, true_and, le_rfl] --- porting note (#10618): simp can prove this --- @[simp] theorem left_mem_Ico : a ∈ Ico a b ↔ a < b := by simp only [mem_Ico, true_and, le_refl] --- porting note (#10618): simp can prove this --- @[simp] theorem right_mem_Icc : b ∈ Icc a b ↔ a ≤ b := by simp only [mem_Icc, and_true, le_rfl] --- porting note (#10618): simp can prove this --- @[simp] theorem right_mem_Ioc : b ∈ Ioc a b ↔ a < b := by simp only [mem_Ioc, and_true, le_rfl] --- porting note (#10618): simp can prove this --- @[simp] theorem left_not_mem_Ioc : a ∉ Ioc a b := fun h => lt_irrefl _ (mem_Ioc.1 h).1 --- porting note (#10618): simp can prove this --- @[simp] theorem left_not_mem_Ioo : a ∉ Ioo a b := fun h => lt_irrefl _ (mem_Ioo.1 h).1 --- porting note (#10618): simp can prove this --- @[simp] theorem right_not_mem_Ico : b ∉ Ico a b := fun h => lt_irrefl _ (mem_Ico.1 h).2 --- porting note (#10618): simp can prove this --- @[simp] theorem right_not_mem_Ioo : b ∉ Ioo a b := fun h => lt_irrefl _ (mem_Ioo.1 h).2 theorem Icc_subset_Icc (ha : a₂ ≤ a₁) (hb : b₁ ≤ b₂) : Icc a₁ b₁ ⊆ Icc a₂ b₂ := by @@ -244,18 +228,12 @@ theorem Icc_ssubset_Icc_right (hI : a₂ ≤ b₂) (ha : a₂ ≤ a₁) (hb : b variable (a) --- porting note (#10618): simp can prove this --- @[simp] theorem Ico_self : Ico a a = ∅ := Ico_eq_empty <| lt_irrefl _ --- porting note (#10618): simp can prove this --- @[simp] theorem Ioc_self : Ioc a a = ∅ := Ioc_eq_empty <| lt_irrefl _ --- porting note (#10618): simp can prove this --- @[simp] theorem Ioo_self : Ioo a a = ∅ := Ioo_eq_empty <| lt_irrefl _ @@ -592,8 +570,6 @@ theorem Ioi_insert [DecidableEq α] (a : α) : insert a (Ioi a) = Ici a := by ext simp_rw [Finset.mem_insert, mem_Ici, mem_Ioi, le_iff_lt_or_eq, or_comm, eq_comm] --- porting note (#10618): simp can prove this --- @[simp] theorem not_mem_Ioi_self {b : α} : b ∉ Ioi b := fun h => lt_irrefl _ (mem_Ioi.1 h) -- Purposefully written the other way around @@ -620,8 +596,6 @@ theorem Iio_insert [DecidableEq α] (b : α) : insert b (Iio b) = Iic b := by ext simp_rw [Finset.mem_insert, mem_Iic, mem_Iio, le_iff_lt_or_eq, or_comm] --- porting note (#10618): simp can prove this --- @[simp] theorem not_mem_Iio_self {b : α} : b ∉ Iio b := fun h => lt_irrefl _ (mem_Iio.1 h) -- Purposefully written the other way around @@ -788,8 +762,6 @@ theorem uIcc_of_ge (h : b ≤ a) : [[a, b]] = Icc b a := by theorem uIcc_comm (a b : α) : [[a, b]] = [[b, a]] := by rw [uIcc, uIcc, inf_comm, sup_comm] --- porting note (#10618): simp can prove this --- @[simp] theorem uIcc_self : [[a, a]] = {a} := by simp [uIcc] @[simp] @@ -802,13 +774,9 @@ theorem Icc_subset_uIcc : Icc a b ⊆ [[a, b]] := theorem Icc_subset_uIcc' : Icc b a ⊆ [[a, b]] := Icc_subset_Icc inf_le_right le_sup_left --- porting note (#10618): simp can prove this --- @[simp] theorem left_mem_uIcc : a ∈ [[a, b]] := mem_Icc.2 ⟨inf_le_left, le_sup_left⟩ --- porting note (#10618): simp can prove this --- @[simp] theorem right_mem_uIcc : b ∈ [[a, b]] := mem_Icc.2 ⟨inf_le_right, le_sup_right⟩ diff --git a/Mathlib/Order/Interval/Finset/Fin.lean b/Mathlib/Order/Interval/Finset/Fin.lean index 24bf50afbdc49..6bf13b1b52462 100644 --- a/Mathlib/Order/Interval/Finset/Fin.lean +++ b/Mathlib/Order/Interval/Finset/Fin.lean @@ -100,23 +100,15 @@ lemma card_Ioo : #(Ioo a b) = b - a - 1 := by rw [← Nat.card_Ioo, ← map_valE theorem card_uIcc : #(uIcc a b) = (b - a : ℤ).natAbs + 1 := by rw [← Nat.card_uIcc, ← map_subtype_embedding_uIcc, card_map] --- Porting note (#10618): simp can prove this --- @[simp] theorem card_fintypeIcc : Fintype.card (Set.Icc a b) = b + 1 - a := by rw [← card_Icc, Fintype.card_ofFinset] --- Porting note (#10618): simp can prove this --- @[simp] theorem card_fintypeIco : Fintype.card (Set.Ico a b) = b - a := by rw [← card_Ico, Fintype.card_ofFinset] --- Porting note (#10618): simp can prove this --- @[simp] theorem card_fintypeIoc : Fintype.card (Set.Ioc a b) = b - a := by rw [← card_Ioc, Fintype.card_ofFinset] --- Porting note (#10618): simp can prove this --- @[simp] theorem card_fintypeIoo : Fintype.card (Set.Ioo a b) = b - a - 1 := by rw [← card_Ioo, Fintype.card_ofFinset] @@ -183,23 +175,15 @@ theorem card_Iic : #(Iic b) = b + 1 := by rw [← Nat.card_Iic b, ← map_valEmb @[simp] theorem card_Iio : #(Iio b) = b := by rw [← Nat.card_Iio b, ← map_valEmbedding_Iio, card_map] --- Porting note (#10618): simp can prove this --- @[simp] theorem card_fintypeIci : Fintype.card (Set.Ici a) = n - a := by rw [Fintype.card_ofFinset, card_Ici] --- Porting note (#10618): simp can prove this --- @[simp] theorem card_fintypeIoi : Fintype.card (Set.Ioi a) = n - 1 - a := by rw [Fintype.card_ofFinset, card_Ioi] --- Porting note (#10618): simp can prove this --- @[simp] theorem card_fintypeIic : Fintype.card (Set.Iic b) = b + 1 := by rw [Fintype.card_ofFinset, card_Iic] --- Porting note (#10618): simp can prove this --- @[simp] theorem card_fintypeIio : Fintype.card (Set.Iio b) = b := by rw [Fintype.card_ofFinset, card_Iio] diff --git a/Mathlib/Order/Interval/Finset/Nat.lean b/Mathlib/Order/Interval/Finset/Nat.lean index 6b580e5787485..e23c82c9f5caa 100644 --- a/Mathlib/Order/Interval/Finset/Nat.lean +++ b/Mathlib/Order/Interval/Finset/Nat.lean @@ -80,33 +80,21 @@ lemma card_Iic : #(Iic b) = b + 1 := by rw [Iic_eq_Icc, card_Icc, Nat.bot_eq_zer @[simp] theorem card_Iio : #(Iio b) = b := by rw [Iio_eq_Ico, card_Ico, Nat.bot_eq_zero, Nat.sub_zero] --- Porting note (#10618): simp can prove this --- @[simp] theorem card_fintypeIcc : Fintype.card (Set.Icc a b) = b + 1 - a := by rw [Fintype.card_ofFinset, card_Icc] --- Porting note (#10618): simp can prove this --- @[simp] theorem card_fintypeIco : Fintype.card (Set.Ico a b) = b - a := by rw [Fintype.card_ofFinset, card_Ico] --- Porting note (#10618): simp can prove this --- @[simp] theorem card_fintypeIoc : Fintype.card (Set.Ioc a b) = b - a := by rw [Fintype.card_ofFinset, card_Ioc] --- Porting note (#10618): simp can prove this --- @[simp] theorem card_fintypeIoo : Fintype.card (Set.Ioo a b) = b - a - 1 := by rw [Fintype.card_ofFinset, card_Ioo] --- Porting note (#10618): simp can prove this --- @[simp] theorem card_fintypeIic : Fintype.card (Set.Iic b) = b + 1 := by rw [Fintype.card_ofFinset, card_Iic] --- Porting note (#10618): simp can prove this --- @[simp] theorem card_fintypeIio : Fintype.card (Set.Iio b) = b := by rw [Fintype.card_ofFinset, card_Iio] -- TODO@Yaël: Generalize all the following lemmas to `SuccOrder` diff --git a/Mathlib/Order/Interval/Multiset.lean b/Mathlib/Order/Interval/Multiset.lean index b8aafdc08f17c..28d6b71306ac8 100644 --- a/Mathlib/Order/Interval/Multiset.lean +++ b/Mathlib/Order/Interval/Multiset.lean @@ -156,13 +156,10 @@ theorem Ioo_eq_zero_of_le (h : b ≤ a) : Ioo a b = 0 := variable (a) --- Porting note (#10618): simp can prove this -- @[simp] theorem Ico_self : Ico a a = 0 := by rw [Ico, Finset.Ico_self, Finset.empty_val] --- Porting note (#10618): simp can prove this -- @[simp] theorem Ioc_self : Ioc a a = 0 := by rw [Ioc, Finset.Ioc_self, Finset.empty_val] --- Porting note (#10618): simp can prove this -- @[simp] theorem Ioo_self : Ioo a a = 0 := by rw [Ioo, Finset.Ioo_self, Finset.empty_val] variable {a} @@ -179,19 +176,15 @@ theorem right_mem_Icc : b ∈ Icc a b ↔ a ≤ b := theorem right_mem_Ioc : b ∈ Ioc a b ↔ a < b := Finset.right_mem_Ioc --- Porting note (#10618): simp can prove this -- @[simp] theorem left_not_mem_Ioc : a ∉ Ioc a b := Finset.left_not_mem_Ioc --- Porting note (#10618): simp can prove this -- @[simp] theorem left_not_mem_Ioo : a ∉ Ioo a b := Finset.left_not_mem_Ioo --- Porting note (#10618): simp can prove this -- @[simp] theorem right_not_mem_Ico : b ∉ Ico a b := Finset.right_not_mem_Ico --- Porting note (#10618): simp can prove this -- @[simp] theorem right_not_mem_Ioo : b ∉ Ioo a b := Finset.right_not_mem_Ioo diff --git a/Mathlib/Order/Interval/Set/Basic.lean b/Mathlib/Order/Interval/Set/Basic.lean index 3dc63cebf960e..8569f24b3fdc5 100644 --- a/Mathlib/Order/Interval/Set/Basic.lean +++ b/Mathlib/Order/Interval/Set/Basic.lean @@ -142,38 +142,22 @@ instance decidableMemIci [Decidable (a ≤ x)] : Decidable (x ∈ Ici a) := by a instance decidableMemIoi [Decidable (a < x)] : Decidable (x ∈ Ioi a) := by assumption --- Porting note (#10618): `simp` can prove this --- @[simp] theorem left_mem_Ioo : a ∈ Ioo a b ↔ False := by simp [lt_irrefl] --- Porting note (#10618): `simp` can prove this --- @[simp] theorem left_mem_Ico : a ∈ Ico a b ↔ a < b := by simp [le_refl] --- Porting note (#10618): `simp` can prove this --- @[simp] theorem left_mem_Icc : a ∈ Icc a b ↔ a ≤ b := by simp [le_refl] --- Porting note (#10618): `simp` can prove this --- @[simp] theorem left_mem_Ioc : a ∈ Ioc a b ↔ False := by simp [lt_irrefl] theorem left_mem_Ici : a ∈ Ici a := by simp --- Porting note (#10618): `simp` can prove this --- @[simp] theorem right_mem_Ioo : b ∈ Ioo a b ↔ False := by simp [lt_irrefl] --- Porting note (#10618): `simp` can prove this --- @[simp] theorem right_mem_Ico : b ∈ Ico a b ↔ False := by simp [lt_irrefl] --- Porting note (#10618): `simp` can prove this --- @[simp] theorem right_mem_Icc : b ∈ Icc a b ↔ a ≤ b := by simp [le_refl] --- Porting note (#10618): `simp` can prove this --- @[simp] theorem right_mem_Ioc : b ∈ Ioc a b ↔ a < b := by simp [le_refl] theorem right_mem_Iic : a ∈ Iic a := by simp @@ -318,18 +302,12 @@ theorem Ioc_eq_empty_of_le (h : b ≤ a) : Ioc a b = ∅ := theorem Ioo_eq_empty_of_le (h : b ≤ a) : Ioo a b = ∅ := Ioo_eq_empty h.not_lt --- Porting note (#10618): `simp` can prove this --- @[simp] theorem Ico_self (a : α) : Ico a a = ∅ := Ico_eq_empty <| lt_irrefl _ --- Porting note (#10618): `simp` can prove this --- @[simp] theorem Ioc_self (a : α) : Ioc a a = ∅ := Ioc_eq_empty <| lt_irrefl _ --- Porting note (#10618): `simp` can prove this --- @[simp] theorem Ioo_self (a : α) : Ioo a a = ∅ := Ioo_eq_empty <| lt_irrefl _ @@ -585,12 +563,8 @@ theorem not_mem_Ico_of_lt (ha : c < a) : c ∉ Ico a b := fun h => ha.not_le h.1 theorem not_mem_Ioc_of_gt (hb : b < c) : c ∉ Ioc a b := fun h => hb.not_le h.2 --- Porting note (#10618): `simp` can prove this --- @[simp] theorem not_mem_Ioi_self : a ∉ Ioi a := lt_irrefl _ --- Porting note (#10618): `simp` can prove this --- @[simp] theorem not_mem_Iio_self : b ∉ Iio b := lt_irrefl _ theorem not_mem_Ioc_of_le (ha : c ≤ a) : c ∉ Ioc a b := fun h => lt_irrefl _ <| h.1.trans_le ha @@ -693,13 +667,9 @@ theorem Ici_diff_Ioi_same : Ici a \ Ioi a = {a} := by theorem Iic_diff_Iio_same : Iic a \ Iio a = {a} := by rw [← Iic_diff_right, diff_diff_cancel_left (singleton_subset_iff.2 right_mem_Iic)] --- Porting note (#10618): `simp` can prove this --- @[simp] theorem Ioi_union_left : Ioi a ∪ {a} = Ici a := ext fun x => by simp [eq_comm, le_iff_eq_or_lt] --- Porting note (#10618): `simp` can prove this --- @[simp] theorem Iio_union_right : Iio a ∪ {a} = Iic a := ext fun _ => le_iff_lt_or_eq.symm diff --git a/Mathlib/Order/Interval/Set/UnorderedInterval.lean b/Mathlib/Order/Interval/Set/UnorderedInterval.lean index ea241a7fbf3dc..22ca451903d02 100644 --- a/Mathlib/Order/Interval/Set/UnorderedInterval.lean +++ b/Mathlib/Order/Interval/Set/UnorderedInterval.lean @@ -74,8 +74,6 @@ lemma uIcc_comm (a b : α) : [[a, b]] = [[b, a]] := by simp_rw [uIcc, inf_comm, lemma uIcc_of_lt (h : a < b) : [[a, b]] = Icc a b := uIcc_of_le h.le lemma uIcc_of_gt (h : b < a) : [[a, b]] = Icc b a := uIcc_of_ge h.le --- Porting note (#10618): `simp` can prove this --- @[simp] lemma uIcc_self : [[a, a]] = {a} := by simp [uIcc] @[simp] lemma nonempty_uIcc : [[a, b]].Nonempty := nonempty_Icc.2 inf_le_sup diff --git a/Mathlib/Order/Irreducible.lean b/Mathlib/Order/Irreducible.lean index 12ebc35f11d80..dc14067f2b2f6 100644 --- a/Mathlib/Order/Irreducible.lean +++ b/Mathlib/Order/Irreducible.lean @@ -161,11 +161,9 @@ theorem InfPrime.inf_le (ha : InfPrime a) : b ⊓ c ≤ a ↔ b ≤ a ∨ c ≤ variable [OrderTop α] {s : Finset ι} {f : ι → α} --- @[simp] Porting note (#10618): simp can prove this. theorem not_infIrred_top : ¬InfIrred (⊤ : α) := isMax_top.not_infIrred --- @[simp] Porting note (#10618): simp can prove this. theorem not_infPrime_top : ¬InfPrime (⊤ : α) := isMax_top.not_infPrime @@ -275,11 +273,9 @@ section LinearOrder variable [LinearOrder α] {a : α} --- @[simp] Porting note (#10618): simp can prove this theorem supPrime_iff_not_isMin : SupPrime a ↔ ¬IsMin a := and_iff_left <| by simp --- @[simp] Porting note (#10618): simp can prove thisrove this theorem infPrime_iff_not_isMax : InfPrime a ↔ ¬IsMax a := and_iff_left <| by simp diff --git a/Mathlib/Order/PartialSups.lean b/Mathlib/Order/PartialSups.lean index 60a016f5f0e41..e40824d604869 100644 --- a/Mathlib/Order/PartialSups.lean +++ b/Mathlib/Order/PartialSups.lean @@ -174,7 +174,6 @@ lemma partialSups_eq_biUnion_range (s : ℕ → Set α) (n : ℕ) : variable [CompleteLattice α] --- Porting note (#10618): simp can prove this @[simp] theorem iSup_partialSups_eq (f : ℕ → α) : ⨆ n, partialSups f n = ⨆ n, f n := ciSup_partialSups_eq <| OrderTop.bddAbove _ From d747c920641e064678ee487e070e58d12045e9a1 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Wed, 23 Oct 2024 19:07:34 +0000 Subject: [PATCH 343/505] feat(AlgebraicGeometry): spreading out morphisms on stalks (#17987) Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- Mathlib.lean | 1 + Mathlib/AlgebraicGeometry/AffineScheme.lean | 25 +- Mathlib/AlgebraicGeometry/SpreadingOut.lean | 337 ++++++++++++++++++++ Mathlib/AlgebraicGeometry/Stalk.lean | 24 ++ 4 files changed, 379 insertions(+), 8 deletions(-) create mode 100644 Mathlib/AlgebraicGeometry/SpreadingOut.lean diff --git a/Mathlib.lean b/Mathlib.lean index 6986c912408c1..8a9b819edad1f 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -906,6 +906,7 @@ import Mathlib.AlgebraicGeometry.Restrict import Mathlib.AlgebraicGeometry.Scheme import Mathlib.AlgebraicGeometry.Sites.BigZariski import Mathlib.AlgebraicGeometry.Spec +import Mathlib.AlgebraicGeometry.SpreadingOut import Mathlib.AlgebraicGeometry.Stalk import Mathlib.AlgebraicGeometry.StructureSheaf import Mathlib.AlgebraicTopology.AlternatingFaceMapComplex diff --git a/Mathlib/AlgebraicGeometry/AffineScheme.lean b/Mathlib/AlgebraicGeometry/AffineScheme.lean index 5a69b75178c41..35e6996d17928 100644 --- a/Mathlib/AlgebraicGeometry/AffineScheme.lean +++ b/Mathlib/AlgebraicGeometry/AffineScheme.lean @@ -248,7 +248,7 @@ theorem iSup_affineOpens_eq_top (X : Scheme) : ⨆ i : X.affineOpens, (i : X.Ope exact isBasis_affine_open X theorem Scheme.map_PrimeSpectrum_basicOpen_of_affine - (X : Scheme) [IsAffine X] (f : Scheme.Γ.obj (op X)) : + (X : Scheme) [IsAffine X] (f : Γ(X, ⊤)) : X.isoSpec.hom ⁻¹ᵁ PrimeSpectrum.basicOpen f = X.basicOpen f := Scheme.toSpecΓ_preimage_basicOpen _ _ @@ -267,23 +267,31 @@ theorem isBasis_basicOpen (X : Scheme) [IsAffine X] : refine ⟨_, ⟨x, rfl⟩, ?_⟩ exact congr_arg Opens.carrier (Scheme.toSpecΓ_preimage_basicOpen _ _).symm +/-- The canonical map `U ⟶ Spec Γ(X, U)` for an open `U ⊆ X`. -/ +noncomputable +def Scheme.Opens.toSpecΓ {X : Scheme.{u}} (U : X.Opens) : + U.toScheme ⟶ Spec Γ(X, U) := + U.toScheme.toSpecΓ ≫ Spec.map U.topIso.inv + namespace IsAffineOpen variable {X Y : Scheme.{u}} {U : X.Opens} (hU : IsAffineOpen U) (f : Γ(X, U)) attribute [-simp] eqToHom_op in /-- The isomorphism `U ≅ Spec Γ(X, U)` for an affine `U`. -/ -@[simps! (config := .lemmasOnly) hom inv] +@[simps! (config := .lemmasOnly) inv] def isoSpec : ↑U ≅ Spec Γ(X, U) := haveI : IsAffine U := hU - U.toScheme.isoSpec ≪≫ Scheme.Spec.mapIso - (X.presheaf.mapIso (eqToIso U.isOpenEmbedding_obj_top).op).op + U.toScheme.isoSpec ≪≫ Scheme.Spec.mapIso U.topIso.symm.op + +lemma isoSpec_hom {X : Scheme.{u}} {U : X.Opens} (hU : IsAffineOpen U) : + hU.isoSpec.hom = U.toSpecΓ := rfl open LocalRing in lemma isoSpec_hom_base_apply (x : U) : - hU.isoSpec.hom.base x = (Spec.map (X.presheaf.germ _ x x.2)).base (closedPoint _) := by - dsimp [IsAffineOpen.isoSpec_hom, Scheme.isoSpec_hom, Scheme.toSpecΓ_base] + hU.isoSpec.hom.base x = (Spec.map (X.presheaf.germ U x x.2)).base (closedPoint _) := by + dsimp [IsAffineOpen.isoSpec_hom, Scheme.isoSpec_hom, Scheme.toSpecΓ_base, Scheme.Opens.toSpecΓ] rw [← Scheme.comp_base_apply, ← Spec.map_comp, (Iso.eq_comp_inv _).mpr (Scheme.Opens.germ_stalkIso_hom U (V := ⊤) x trivial), X.presheaf.germ_res_assoc, Spec.map_comp, Scheme.comp_base_apply] @@ -329,7 +337,7 @@ theorem range_fromSpec : infer_instance @[simp] -theorem opensRange_fromSpec : Scheme.Hom.opensRange hU.fromSpec = U := Opens.ext (range_fromSpec hU) +theorem opensRange_fromSpec :hU.fromSpec.opensRange = U := Opens.ext (range_fromSpec hU) @[reassoc (attr := simp)] theorem map_fromSpec {V : X.Opens} (hV : IsAffineOpen V) (f : op U ⟶ op V) : @@ -349,7 +357,8 @@ lemma Spec_map_appLE_fromSpec (f : X ⟶ Y) {V : X.Opens} {U : Y.Opens} Spec.map (f.appLE U V i) ≫ hU.fromSpec = hV.fromSpec ≫ f := by have : IsAffine U := hU simp only [IsAffineOpen.fromSpec, Category.assoc, isoSpec_inv] - rw [← Scheme.restrictFunctor_map_ofRestrict (homOfLE i), Category.assoc, ← morphismRestrict_ι, + simp_rw [← Scheme.restrictFunctor_map_ofRestrict (homOfLE i)] + rw [Category.assoc, ← morphismRestrict_ι, ← Category.assoc _ (f ∣_ U) U.ι, ← @Scheme.isoSpec_inv_naturality_assoc, ← Spec.map_comp_assoc, ← Spec.map_comp_assoc, Scheme.comp_app, morphismRestrict_app, Scheme.restrictFunctor_map_app, Scheme.Hom.app_eq_appLE, Scheme.Hom.appLE_map, diff --git a/Mathlib/AlgebraicGeometry/SpreadingOut.lean b/Mathlib/AlgebraicGeometry/SpreadingOut.lean new file mode 100644 index 0000000000000..e5ee6354c8f68 --- /dev/null +++ b/Mathlib/AlgebraicGeometry/SpreadingOut.lean @@ -0,0 +1,337 @@ +/- +Copyright (c) 2024 Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Andrew Yang +-/ +import Mathlib.AlgebraicGeometry.Morphisms.FiniteType +import Mathlib.AlgebraicGeometry.Noetherian +import Mathlib.AlgebraicGeometry.Stalk +import Mathlib.AlgebraicGeometry.Properties + +/-! +# Spreading out morphisms + +Under certain conditions, a morphism on stalks `Spec 𝒪_{X, x} ⟶ Spec 𝒪_{Y, y}` can be spread +out into a neighborhood of `x`. + +## Main result +Given `S`-schemes `X Y` and points `x : X` `y : Y` over `s : S`. +Suppose we have the following diagram of `S`-schemes +``` +Spec 𝒪_{X, x} ⟶ X + | + Spec(φ) + ↓ +Spec 𝒪_{Y, y} ⟶ Y +``` +We would like to spread `Spec(φ)` out to an `S`-morphism on an open subscheme `U ⊆ X` +``` +Spec 𝒪_{X, x} ⟶ U ⊆ X + | | + Spec(φ) | + ↓ ↓ +Spec 𝒪_{Y, y} ⟶ Y +``` +- `AlgebraicGeometry.spread_out_unique_of_isGermInjective`: + The lift is "unique" if the germ map is injective at `x`. +- `AlgebraicGeometry.spread_out_of_isGermInjective`: + The lift exists if `Y` is locally of finite type and the germ map is injective at `x`. + +## TODO + +Show that certain morphism properties can also be spread out. + +-/ + +universe u + +open CategoryTheory + +namespace AlgebraicGeometry + +variable {X Y S : Scheme.{u}} (f : X ⟶ Y) (sX : X ⟶ S) (sY : Y ⟶ S) (e : f ≫ sY = sX) +variable {R A : CommRingCat.{u}} + +/-- The germ map at `x` is injective if there exists some affine `U ∋ x` + such that the map `Γ(X, U) ⟶ X_x` is injective -/ +class Scheme.IsGermInjectiveAt (X : Scheme.{u}) (x : X) : Prop where + cond : ∃ (U : X.Opens) (hx : x ∈ U), IsAffineOpen U ∧ Function.Injective (X.presheaf.germ U x hx) + +lemma injective_germ_basicOpen (U : X.Opens) (hU : IsAffineOpen U) + (x : X) (hx : x ∈ U) (f : Γ(X, U)) + (hf : x ∈ X.basicOpen f) + (H : Function.Injective (X.presheaf.germ U x hx)) : + Function.Injective (X.presheaf.germ (X.basicOpen f) x hf) := by + rw [RingHom.injective_iff_ker_eq_bot, RingHom.ker_eq_bot_iff_eq_zero] at H ⊢ + intros t ht + have := hU.isLocalization_basicOpen f + obtain ⟨t, s, rfl⟩ := IsLocalization.mk'_surjective (.powers f) t + rw [← RingHom.mem_ker, IsLocalization.mk'_eq_mul_mk'_one, Ideal.mul_unit_mem_iff_mem, + RingHom.mem_ker, RingHom.algebraMap_toAlgebra, X.presheaf.germ_res_apply] at ht + swap; · exact @isUnit_of_invertible _ _ _ (@IsLocalization.invertible_mk'_one ..) + rw [H _ ht, IsLocalization.mk'_zero] + +lemma Scheme.exists_germ_injective (X : Scheme.{u}) (x : X) [X.IsGermInjectiveAt x] : + ∃ (U : X.Opens) (hx : x ∈ U), + IsAffineOpen U ∧ Function.Injective (X.presheaf.germ U x hx) := + Scheme.IsGermInjectiveAt.cond + +lemma Scheme.exists_le_and_germ_injective (X : Scheme.{u}) (x : X) [X.IsGermInjectiveAt x] + (V : X.Opens) (hxV : x ∈ V) : + ∃ (U : X.Opens) (hx : x ∈ U), + IsAffineOpen U ∧ U ≤ V ∧ Function.Injective (X.presheaf.germ U x hx) := by + obtain ⟨U, hx, hU, H⟩ := Scheme.IsGermInjectiveAt.cond (x := x) + obtain ⟨f, hf, hxf⟩ := hU.exists_basicOpen_le ⟨x, hxV⟩ hx + exact ⟨X.basicOpen f, hxf, hU.basicOpen f, hf, injective_germ_basicOpen U hU x hx f hxf H⟩ + +instance (x : X) [X.IsGermInjectiveAt x] [IsOpenImmersion f] : + Y.IsGermInjectiveAt (f.base x) := by + obtain ⟨U, hxU, hU, H⟩ := X.exists_germ_injective x + refine ⟨⟨f ''ᵁ U, ⟨x, hxU, rfl⟩, hU.image_of_isOpenImmersion f, ?_⟩⟩ + refine ((MorphismProperty.injective CommRingCat).cancel_right_of_respectsIso _ + (f.stalkMap x)).mp ?_ + refine ((MorphismProperty.injective CommRingCat).cancel_left_of_respectsIso + (f.appIso U).inv _).mp ?_ + simpa + +variable {f} in +lemma isGermInjectiveAt_iff_of_isOpenImmersion {x : X} [IsOpenImmersion f]: + Y.IsGermInjectiveAt (f.base x) ↔ X.IsGermInjectiveAt x := by + refine ⟨fun H ↦ ?_, fun _ ↦ inferInstance⟩ + obtain ⟨U, hxU, hU, hU', H⟩ := + Y.exists_le_and_germ_injective (f.base x) (V := f.opensRange) ⟨x, rfl⟩ + obtain ⟨V, hV⟩ := (IsOpenImmersion.affineOpensEquiv f).surjective ⟨⟨U, hU⟩, hU'⟩ + obtain rfl : f ''ᵁ V = U := Subtype.eq_iff.mp (Subtype.eq_iff.mp hV) + obtain ⟨y, hy, e : f.base y = f.base x⟩ := hxU + obtain rfl := f.isOpenEmbedding.inj e + refine ⟨V, hy, V.2, ?_⟩ + replace H := ((MorphismProperty.injective CommRingCat).cancel_right_of_respectsIso _ + (f.stalkMap y)).mpr H + replace H := ((MorphismProperty.injective CommRingCat).cancel_left_of_respectsIso + (f.appIso V).inv _).mpr H + simpa using H + +/-- +The class of schemes such that for each `x : X`, +`Γ(X, U) ⟶ X_x` is injective for some affine `U` containing `x`. + +This is typically satisfied when `X` is integral or locally noetherian. +-/ +abbrev Scheme.IsGermInjective (X : Scheme.{u}) := ∀ x : X, X.IsGermInjectiveAt x + +lemma Scheme.IsGermInjective.of_openCover + {X : Scheme.{u}} (𝒰 : X.OpenCover) [∀ i, (𝒰.obj i).IsGermInjective] : X.IsGermInjective := by + intro x + rw [← (𝒰.covers x).choose_spec] + infer_instance + +protected +lemma Scheme.IsGermInjective.Spec + (H : ∀ I : Ideal R, I.IsPrime → ∃ f : R, f ∉ I ∧ ∀ (x y : R) + (_ : y * x = 0) (_ : y ∉ I), ∃ n, f ^ n * x = 0) : (Spec R).IsGermInjective := by + refine fun p ↦ ⟨?_⟩ + obtain ⟨f, hf, H⟩ := H p.asIdeal p.2 + refine ⟨PrimeSpectrum.basicOpen f, hf, ?_, ?_⟩ + · rw [← basicOpen_eq_of_affine] + exact (isAffineOpen_top (Spec R)).basicOpen _ + rw [RingHom.injective_iff_ker_eq_bot, RingHom.ker_eq_bot_iff_eq_zero] + intro x hx + obtain ⟨x, s, rfl⟩ := IsLocalization.mk'_surjective + (S := ((Spec.structureSheaf R).val.obj (.op <| PrimeSpectrum.basicOpen f))) (.powers f) x + rw [← RingHom.mem_ker, IsLocalization.mk'_eq_mul_mk'_one, Ideal.mul_unit_mem_iff_mem, + RingHom.mem_ker, RingHom.algebraMap_toAlgebra] at hx + swap; · exact @isUnit_of_invertible _ _ _ (@IsLocalization.invertible_mk'_one ..) + erw [StructureSheaf.germ_toOpen] at hx + obtain ⟨⟨y, hy⟩, hy'⟩ := (IsLocalization.map_eq_zero_iff p.asIdeal.primeCompl + ((Spec.structureSheaf R).presheaf.stalk p) _).mp hx + obtain ⟨n, hn⟩ := H x y hy' hy + refine (@IsLocalization.mk'_eq_zero_iff ..).mpr ?_ + exact ⟨⟨_, n, rfl⟩, hn⟩ + +instance (priority := 100) [IsIntegral X] : X.IsGermInjective := by + refine fun x ↦ ⟨⟨(X.affineCover.map x).opensRange, X.affineCover.covers x, + (isAffineOpen_opensRange (X.affineCover.map x)), ?_⟩⟩ + have : Nonempty (X.affineCover.map x).opensRange := ⟨⟨_, X.affineCover.covers x⟩⟩ + have := (isAffineOpen_opensRange (X.affineCover.map x)).isLocalization_stalk + ⟨_, X.affineCover.covers x⟩ + exact @IsLocalization.injective _ _ _ _ _ (show _ from _) this + (Ideal.primeCompl_le_nonZeroDivisors _) + +instance (priority := 100) [IsLocallyNoetherian X] : X.IsGermInjective := by + suffices ∀ (R : CommRingCat.{u}) (_ : IsNoetherianRing R), (Spec R).IsGermInjective by + refine @Scheme.IsGermInjective.of_openCover _ (X.affineOpenCover.openCover) (fun i ↦ this _ ?_) + have := isLocallyNoetherian_of_isOpenImmersion (X.affineOpenCover.map i) + infer_instance + refine fun R hR ↦ Scheme.IsGermInjective.Spec fun I hI ↦ ?_ + let J := RingHom.ker <| algebraMap R (Localization.AtPrime I) + have hJ (x) : x ∈ J ↔ ∃ y : I.primeCompl, y * x = 0 := + IsLocalization.map_eq_zero_iff I.primeCompl _ x + choose f hf using fun x ↦ (hJ x).mp + obtain ⟨s, hs⟩ := (isNoetherianRing_iff_ideal_fg R).mp ‹_› J + have hs' : (s : Set R) ⊆ J := hs ▸ Ideal.subset_span + refine ⟨_, (s.attach.prod fun x ↦ f x (hs' x.2)).2, fun x y e hy ↦ ⟨1, ?_⟩⟩ + rw [pow_one, mul_comm, ← smul_eq_mul, ← Submodule.mem_annihilator_span_singleton] + refine SetLike.le_def.mp ?_ ((hJ x).mpr ⟨⟨y, hy⟩, e⟩) + rw [← hs, Ideal.span_le] + intro i hi + rw [SetLike.mem_coe, Submodule.mem_annihilator_span_singleton, smul_eq_mul, + mul_comm, ← smul_eq_mul, ← Submodule.mem_annihilator_span_singleton, Submonoid.coe_finset_prod] + refine Ideal.mem_of_dvd _ (Finset.dvd_prod_of_mem _ (s.mem_attach ⟨i, hi⟩)) ?_ + rw [Submodule.mem_annihilator_span_singleton, smul_eq_mul] + exact hf i _ + +/-- +Let `x : X` and `f g : X ⟶ Y` be two morphisms such that `f x = g x`. +If `f` and `g` agree on the stalk of `x`, then they agree on an open neighborhood of `x`, +provided `X` is "germ-injective" at `x` (e.g. when it's integral or locally noetherian). + +TODO: The condition on `X` is unnecessary when `Y` is locally of finite type. +-/ +@[stacks 0BX6] +lemma spread_out_unique_of_isGermInjective {x : X} [X.IsGermInjectiveAt x] + (f g : X ⟶ Y) (e : f.base x = g.base x) + (H : f.stalkMap x = + Y.presheaf.stalkSpecializes (Inseparable.of_eq e.symm).specializes ≫ g.stalkMap x) : + ∃ (U : X.Opens), x ∈ U ∧ U.ι ≫ f = U.ι ≫ g := by + obtain ⟨_, ⟨V : Y.Opens, hV, rfl⟩, hxV, -⟩ := + (isBasis_affine_open Y).exists_subset_of_mem_open (Set.mem_univ (f.base x)) isOpen_univ + have hxV' : g.base x ∈ V := e ▸ hxV + obtain ⟨U, hxU, _, hUV, HU⟩ := X.exists_le_and_germ_injective x (f ⁻¹ᵁ V ⊓ g ⁻¹ᵁ V) ⟨hxV, hxV'⟩ + refine ⟨U, hxU, ?_⟩ + rw [← Scheme.Hom.resLE_comp_ι _ (hUV.trans inf_le_left), + ← Scheme.Hom.resLE_comp_ι _ (hUV.trans inf_le_right)] + congr 1 + have : IsAffine V := hV + suffices ∀ (U₀ V₀) (eU : U = U₀) (eV : V = V₀), + f.appLE V₀ U₀ (eU ▸ eV ▸ hUV.trans inf_le_left) = + g.appLE V₀ U₀ (eU ▸ eV ▸ hUV.trans inf_le_right) by + rw [← cancel_mono V.toScheme.isoSpec.hom] + simp only [Scheme.isoSpec, asIso_hom, Scheme.toSpecΓ_naturality, + Scheme.Hom.app_eq_appLE, Scheme.Hom.resLE_appLE] + congr 2 + apply this <;> simp + rintro U V rfl rfl + have := ConcreteCategory.mono_of_injective _ HU + rw [← cancel_mono (X.presheaf.germ U x hxU)] + simp only [Scheme.Hom.appLE, Category.assoc, X.presheaf.germ_res', ← Scheme.stalkMap_germ, H] + simp only [TopCat.Presheaf.germ_stalkSpecializes_assoc, Scheme.stalkMap_germ] + +lemma exists_lift_of_germInjective_aux {U : X.Opens} {x : X} (hxU) + (φ : A ⟶ X.presheaf.stalk x) (φRA : R ⟶ A) (φRX : R ⟶ Γ(X, U)) + (hφRA : RingHom.FiniteType φRA) + (e : φRA ≫ φ = φRX ≫ X.presheaf.germ U x hxU) : + ∃ (V : X.Opens) (hxV : x ∈ V), + V ≤ U ∧ RingHom.range φ ≤ RingHom.range (X.presheaf.germ V x hxV) := by + letI := φRA.toAlgebra + obtain ⟨s, hs⟩ := hφRA + choose W hxW f hf using fun t ↦ X.presheaf.germ_exist x (φ t) + have H : x ∈ s.inf W ⊓ U := by + rw [← SetLike.mem_coe, TopologicalSpace.Opens.coe_inf, TopologicalSpace.Opens.coe_finset_inf] + exact ⟨by simpa using fun x _ ↦ hxW x, hxU⟩ + refine ⟨s.inf W ⊓ U, H, inf_le_right, ?_⟩ + letI := φRX.toAlgebra + letI := (φRX ≫ X.presheaf.germ U x hxU).toAlgebra + letI := (φRX ≫ X.presheaf.map (homOfLE (inf_le_right (a := s.inf W))).op).toAlgebra + let φ' : A →ₐ[R] X.presheaf.stalk x := { φ with commutes' := DFunLike.congr_fun e } + let ψ : Γ(X, s.inf W ⊓ U) →ₐ[R] X.presheaf.stalk x := + { X.presheaf.germ _ x H with commutes' := fun x ↦ X.presheaf.germ_res_apply _ _ _ _ } + show AlgHom.range φ' ≤ AlgHom.range ψ + rw [← Algebra.map_top, ← hs, AlgHom.map_adjoin, Algebra.adjoin_le_iff] + rintro _ ⟨i, hi, rfl : φ i = _⟩ + refine ⟨X.presheaf.map (homOfLE (inf_le_left.trans (Finset.inf_le hi))).op (f i), ?_⟩ + exact (X.presheaf.germ_res_apply _ _ _ _).trans (hf _) + +/-- +Suppose `X` is a scheme, `x : X` such that the germ map at `x` is (locally) injective, +and `U` is a neighborhood of `x`. +Given a commutative diagram of `CommRingCat` +``` +R ⟶ Γ(X, U) +↓ ↓ +A ⟶ 𝒪_{X, x} +``` +such that `R` is of finite type over `A`, we may lift `A ⟶ 𝒪_{X, x}` to some `A ⟶ Γ(X, V)`. +-/ +lemma exists_lift_of_germInjective {x : X} [X.IsGermInjectiveAt x] {U : X.Opens} (hxU : x ∈ U) + (φ : A ⟶ X.presheaf.stalk x) (φRA : R ⟶ A) (φRX : R ⟶ Γ(X, U)) + (hφRA : RingHom.FiniteType φRA) + (e : φRA ≫ φ = φRX ≫ X.presheaf.germ U x hxU) : + ∃ (V : X.Opens) (hxV : x ∈ V) (φ' : A ⟶ Γ(X, V)) (i : V ≤ U), IsAffineOpen V ∧ + φ = φ' ≫ X.presheaf.germ V x hxV ∧ φRX ≫ X.presheaf.map i.hom.op = φRA ≫ φ' := by + obtain ⟨V, hxV, iVU, hV⟩ := exists_lift_of_germInjective_aux hxU φ φRA φRX hφRA e + obtain ⟨V', hxV', hV', iV'V, H⟩ := X.exists_le_and_germ_injective x V hxV + let f := X.presheaf.germ V' x hxV' + have hf' : RingHom.range (X.presheaf.germ V x hxV) ≤ RingHom.range f := by + rw [← X.presheaf.germ_res iV'V.hom _ hxV'] + exact Set.range_comp_subset_range (X.presheaf.map iV'V.hom.op) f + let e := RingEquiv.ofLeftInverse H.hasLeftInverse.choose_spec + refine ⟨V', hxV', CommRingCat.ofHom (e.symm.toRingHom.comp + (φ.codRestrict _ (fun x ↦ hf' (hV ⟨x, rfl⟩)))), iV'V.trans iVU, hV', ?_, ?_⟩ + · ext a + show φ a = (e (e.symm _)).1 + simp only [RingEquiv.apply_symm_apply] + rfl + · ext a + apply e.injective + show e _ = e (e.symm _) + rw [RingEquiv.apply_symm_apply] + ext + show X.presheaf.germ _ _ _ (X.presheaf.map _ _) = (φRA ≫ φ) a + rw [X.presheaf.germ_res_apply, ‹φRA ≫ φ = _›] + rfl + +/-- +Given `S`-schemes `X Y` and points `x : X` `y : Y` over `s : S`. +Suppose we have the following diagram of `S`-schemes +``` +Spec 𝒪_{X, x} ⟶ X + | + Spec(φ) + ↓ +Spec 𝒪_{Y, y} ⟶ Y +``` +Then the map `Spec(φ)` spreads out to an `S`-morphism on an open subscheme `U ⊆ X`, +``` +Spec 𝒪_{X, x} ⟶ U ⊆ X + | | + Spec(φ) | + ↓ ↓ +Spec 𝒪_{Y, y} ⟶ Y +``` +provided that `Y` is locally of finite type over `S` and +`X` is "germ-injective" at `x` (e.g. when it's integral or locally noetherian). + +TODO: The condition on `X` is unnecessary when `Y` is locally of finite presentation. +-/ +@[stacks 0BX6] +lemma spread_out_of_isGermInjective [LocallyOfFiniteType sY] {x : X} [X.IsGermInjectiveAt x] {y : Y} + (e : sX.base x = sY.base y) (φ : Y.presheaf.stalk y ⟶ X.presheaf.stalk x) + (h : sY.stalkMap y ≫ φ = + S.presheaf.stalkSpecializes (Inseparable.of_eq e).specializes ≫ sX.stalkMap x) : + ∃ (U : X.Opens) (hxU : x ∈ U) (f : U.toScheme ⟶ Y), + Spec.map φ ≫ Y.fromSpecStalk y = U.fromSpecStalkOfMem x hxU ≫ f ∧ + f ≫ sY = U.ι ≫ sX := by + obtain ⟨_, ⟨U, hU, rfl⟩, hxU, -⟩ := + (isBasis_affine_open S).exists_subset_of_mem_open (Set.mem_univ (sX.base x)) isOpen_univ + have hyU : sY.base y ∈ U := e ▸ hxU + obtain ⟨_, ⟨V : Y.Opens, hV, rfl⟩, hyV, iVU⟩ := + (isBasis_affine_open Y).exists_subset_of_mem_open hyU (sY ⁻¹ᵁ U).2 + have : sY.appLE U V iVU ≫ Y.presheaf.germ V y hyV ≫ φ = + sX.app U ≫ X.presheaf.germ (sX ⁻¹ᵁ U) x hxU := by + rw [Scheme.Hom.appLE, Category.assoc, Y.presheaf.germ_res_assoc, + ← Scheme.stalkMap_germ_assoc, h] + simp + obtain ⟨W, hxW, φ', i, hW, h₁, h₂⟩ := + exists_lift_of_germInjective (R := Γ(S, U)) (A := Γ(Y, V)) (U := sX ⁻¹ᵁ U) (x := x) hxU + (Y.presheaf.germ _ y hyV ≫ φ) (sY.appLE U V iVU) (sX.app U) + (LocallyOfFiniteType.finiteType_of_affine_subset ⟨_, hU⟩ ⟨_, hV⟩ _) this + refine ⟨W, hxW, W.toSpecΓ ≫ Spec.map φ' ≫ hV.fromSpec, ?_, ?_⟩ + · rw [W.fromSpecStalkOfMem_toSpecΓ_assoc x hxW, ← Spec.map_comp_assoc, ← h₁, + Spec.map_comp, Category.assoc, ← IsAffineOpen.fromSpecStalk, + IsAffineOpen.fromSpecStalk_eq_fromSpecStalk] + · simp only [Category.assoc, IsAffineOpen.isoSpec_inv_ι_assoc] + rw [← IsAffineOpen.Spec_map_appLE_fromSpec sY hU hV iVU, ← Spec.map_comp_assoc, ← h₂, + ← Scheme.Hom.appLE, ← hW.isoSpec_hom, IsAffineOpen.Spec_map_appLE_fromSpec sX hU hW i, + ← Iso.eq_inv_comp, IsAffineOpen.isoSpec_inv_ι_assoc] + +end AlgebraicGeometry diff --git a/Mathlib/AlgebraicGeometry/Stalk.lean b/Mathlib/AlgebraicGeometry/Stalk.lean index 92b56c3849cc3..b5c850490ad3c 100644 --- a/Mathlib/AlgebraicGeometry/Stalk.lean +++ b/Mathlib/AlgebraicGeometry/Stalk.lean @@ -155,6 +155,30 @@ lemma range_fromSpecStalk {x : X} : rw [← Spec_map_stalkSpecializes_fromSpecStalk hy] at this exact ⟨_, this⟩ +/-- The canonical map `Spec 𝒪_{X, x} ⟶ U` given `x ∈ U ⊆ X`. -/ +noncomputable +def Opens.fromSpecStalkOfMem {X : Scheme.{u}} (U : X.Opens) (x : X) (hxU : x ∈ U) : + Spec (X.presheaf.stalk x) ⟶ U := + Spec.map (inv (U.ι.stalkMap ⟨x, hxU⟩)) ≫ U.toScheme.fromSpecStalk ⟨x, hxU⟩ + +@[reassoc] +lemma fromSpecStalk_toSpecΓ (X : Scheme.{u}) (x : X) : + X.fromSpecStalk x ≫ X.toSpecΓ = Spec.map (X.presheaf.germ ⊤ x trivial) := by + rw [Scheme.toSpecΓ_naturality, ← SpecMap_ΓSpecIso_hom, ← Spec.map_comp, + @Scheme.fromSpecStalk_app X ⊤ _ trivial] + simp + +@[reassoc (attr := simp)] +lemma Opens.fromSpecStalkOfMem_toSpecΓ {X : Scheme.{u}} (U : X.Opens) (x : X) (hxU : x ∈ U) : + U.fromSpecStalkOfMem x hxU ≫ U.toSpecΓ = Spec.map (X.presheaf.germ U x hxU) := by + rw [fromSpecStalkOfMem, Opens.toSpecΓ, Category.assoc, fromSpecStalk_toSpecΓ_assoc, + ← Spec.map_comp, ← Spec.map_comp] + congr 1 + rw [IsIso.comp_inv_eq, Iso.inv_comp_eq] + erw [stalkMap_germ U.ι U ⟨x, hxU⟩] + rw [Opens.ι_app, Opens.topIso_hom, ← Functor.map_comp_assoc] + exact (U.toScheme.presheaf.germ_res (homOfLE le_top) ⟨x, hxU⟩ (U := U.ι ⁻¹ᵁ U) hxU).symm + end Scheme end fromSpecStalk From a6c3c5234f201d77e6f9268325462db480917d4d Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 23 Oct 2024 19:30:45 +0000 Subject: [PATCH 344/505] =?UTF-8?q?feat:=20`f=20a=20=E2=89=A4=20f=20(succ?= =?UTF-8?q?=20a)=20=E2=86=92=20Monotone=20f`=20(#18008)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit From GrowthInGroups --- Mathlib/Algebra/Order/SuccPred.lean | 72 +++++++++++++- Mathlib/Order/Interval/Set/Monotone.lean | 42 +-------- Mathlib/Order/SuccPred/Archimedean.lean | 114 +++++++++++++++++++++++ 3 files changed, 188 insertions(+), 40 deletions(-) diff --git a/Mathlib/Algebra/Order/SuccPred.lean b/Mathlib/Algebra/Order/SuccPred.lean index 9ef68870ecd16..185ea8250d6f9 100644 --- a/Mathlib/Algebra/Order/SuccPred.lean +++ b/Mathlib/Algebra/Order/SuccPred.lean @@ -6,7 +6,7 @@ Authors: Violeta Hernández Palacios, Yaël Dillies import Mathlib.Algebra.Group.Basic import Mathlib.Algebra.Order.ZeroLEOne import Mathlib.Data.Int.Cast.Defs -import Mathlib.Order.SuccPred.Basic +import Mathlib.Order.SuccPred.Archimedean /-! # Interaction between successors and arithmetic @@ -183,3 +183,73 @@ theorem lt_one_iff_nonpos [AddMonoidWithOne α] [ZeroLEOneClass α] [NeZero (1 : end LinearOrder end Order + +section Monotone +variable {α β : Type*} [PartialOrder α] [Preorder β] + +section SuccAddOrder +variable [Add α] [One α] [SuccAddOrder α] [IsSuccArchimedean α] {s : Set α} {f : α → β} + +lemma monotoneOn_of_le_add_one (hs : s.OrdConnected) : + (∀ a, ¬ IsMax a → a ∈ s → a + 1 ∈ s → f a ≤ f (a + 1)) → MonotoneOn f s := by + simpa [Order.succ_eq_add_one] using monotoneOn_of_le_succ hs (f := f) + +lemma antitoneOn_of_add_one_le (hs : s.OrdConnected) : + (∀ a, ¬ IsMax a → a ∈ s → a + 1 ∈ s → f (a + 1) ≤ f a) → AntitoneOn f s := by + simpa [Order.succ_eq_add_one] using antitoneOn_of_succ_le hs (f := f) + +lemma strictMonoOn_of_lt_add_one (hs : s.OrdConnected) : + (∀ a, ¬ IsMax a → a ∈ s → a + 1 ∈ s → f a < f (a + 1)) → StrictMonoOn f s := by + simpa [Order.succ_eq_add_one] using strictMonoOn_of_lt_succ hs (f := f) + +lemma strictAntiOn_of_add_one_lt (hs : s.OrdConnected) : + (∀ a, ¬ IsMax a → a ∈ s → a + 1 ∈ s → f (a + 1) < f a) → StrictAntiOn f s := by + simpa [Order.succ_eq_add_one] using strictAntiOn_of_succ_lt hs (f := f) + +lemma monotone_of_le_add_one : (∀ a, ¬ IsMax a → f a ≤ f (a + 1)) → Monotone f := by + simpa [Order.succ_eq_add_one] using monotone_of_le_succ (f := f) + +lemma antitone_of_add_one_le : (∀ a, ¬ IsMax a → f (a + 1) ≤ f a) → Antitone f := by + simpa [Order.succ_eq_add_one] using antitone_of_succ_le (f := f) + +lemma strictMono_of_lt_add_one : (∀ a, ¬ IsMax a → f a < f (a + 1)) → StrictMono f := by + simpa [Order.succ_eq_add_one] using strictMono_of_lt_succ (f := f) + +lemma strictAnti_of_add_one_lt : (∀ a, ¬ IsMax a → f (a + 1) < f a) → StrictAnti f := by + simpa [Order.succ_eq_add_one] using strictAnti_of_succ_lt (f := f) + +end SuccAddOrder + +section PredSubOrder +variable [Sub α] [One α] [PredSubOrder α] [IsPredArchimedean α] {s : Set α} {f : α → β} + +lemma monotoneOn_of_sub_one_le (hs : s.OrdConnected) : + (∀ a, ¬ IsMin a → a ∈ s → a - 1 ∈ s → f (a - 1) ≤ f a) → MonotoneOn f s := by + simpa [Order.pred_eq_sub_one] using monotoneOn_of_pred_le hs (f := f) + +lemma antitoneOn_of_le_sub_one (hs : s.OrdConnected) : + (∀ a, ¬ IsMin a → a ∈ s → a - 1 ∈ s → f a ≤ f (a - 1)) → AntitoneOn f s := by + simpa [Order.pred_eq_sub_one] using antitoneOn_of_le_pred hs (f := f) + +lemma strictMonoOn_of_sub_one_lt (hs : s.OrdConnected) : + (∀ a, ¬ IsMin a → a ∈ s → a - 1 ∈ s → f (a - 1) < f a) → StrictMonoOn f s := by + simpa [Order.pred_eq_sub_one] using strictMonoOn_of_pred_lt hs (f := f) + +lemma strictAntiOn_of_lt_sub_one (hs : s.OrdConnected) : + (∀ a, ¬ IsMin a → a ∈ s → a - 1 ∈ s → f a < f (a - 1)) → StrictAntiOn f s := by + simpa [Order.pred_eq_sub_one] using strictAntiOn_of_lt_pred hs (f := f) + +lemma monotone_of_sub_one_le : (∀ a, ¬ IsMin a → f (a - 1) ≤ f a) → Monotone f := by + simpa [Order.pred_eq_sub_one] using monotone_of_pred_le (f := f) + +lemma antitone_of_le_sub_one : (∀ a, ¬ IsMin a → f a ≤ f (a - 1)) → Antitone f := by + simpa [Order.pred_eq_sub_one] using antitone_of_le_pred (f := f) + +lemma strictMono_of_sub_one_lt : (∀ a, ¬ IsMin a → f (a - 1) < f a) → StrictMono f := by + simpa [Order.pred_eq_sub_one] using strictMono_of_pred_lt (f := f) + +lemma strictAnti_of_lt_sub_one : (∀ a, ¬ IsMin a → f a < f (a - 1)) → StrictAnti f := by + simpa [Order.pred_eq_sub_one] using strictAnti_of_lt_pred (f := f) + +end PredSubOrder +end Monotone diff --git a/Mathlib/Order/Interval/Set/Monotone.lean b/Mathlib/Order/Interval/Set/Monotone.lean index 6b8430931f6c7..f2f3f914fb9c1 100644 --- a/Mathlib/Order/Interval/Set/Monotone.lean +++ b/Mathlib/Order/Interval/Set/Monotone.lean @@ -164,58 +164,22 @@ variable {α β : Type*} [PartialOrder α] [Preorder β] {ψ : α → β} /-- A function `ψ` on a `SuccOrder` is strictly monotone before some `n` if for all `m` such that `m < n`, we have `ψ m < ψ (succ m)`. -/ theorem strictMonoOn_Iic_of_lt_succ [SuccOrder α] [IsSuccArchimedean α] {n : α} - (hψ : ∀ m, m < n → ψ m < ψ (succ m)) : StrictMonoOn ψ (Set.Iic n) := by - intro x hx y hy hxy - obtain ⟨i, rfl⟩ := hxy.le.exists_succ_iterate - induction' i with k ih - · simp at hxy - cases' k with k - · exact hψ _ (lt_of_lt_of_le hxy hy) - rw [Set.mem_Iic] at * - simp only [Function.iterate_succ', Function.comp_apply] at ih hxy hy ⊢ - by_cases hmax : IsMax (succ^[k] x) - · rw [succ_eq_iff_isMax.2 hmax] at hxy ⊢ - exact ih (le_trans (le_succ _) hy) hxy - by_cases hmax' : IsMax (succ (succ^[k] x)) - · rw [succ_eq_iff_isMax.2 hmax'] at hxy ⊢ - exact ih (le_trans (le_succ _) hy) hxy - refine - lt_trans - (ih (le_trans (le_succ _) hy) - (lt_of_le_of_lt (le_succ_iterate k _) (lt_succ_iff_not_isMax.2 hmax))) - ?_ - rw [← Function.comp_apply (f := succ), ← Function.iterate_succ'] - refine hψ _ (lt_of_lt_of_le ?_ hy) - rwa [Function.iterate_succ', Function.comp_apply, lt_succ_iff_not_isMax] - -theorem strictMono_of_lt_succ [SuccOrder α] [IsSuccArchimedean α] - (hψ : ∀ m, ψ m < ψ (succ m)) : StrictMono ψ := fun _ _ h ↦ - (strictMonoOn_Iic_of_lt_succ fun m _ ↦ hψ m) h.le le_rfl h + (hψ : ∀ m, m < n → ψ m < ψ (succ m)) : StrictMonoOn ψ (Set.Iic n) := + strictMonoOn_of_lt_succ ordConnected_Iic fun _a ha' _ ha ↦ + hψ _ <| (succ_le_iff_of_not_isMax ha').1 ha theorem strictAntiOn_Iic_of_succ_lt [SuccOrder α] [IsSuccArchimedean α] {n : α} (hψ : ∀ m, m < n → ψ (succ m) < ψ m) : StrictAntiOn ψ (Set.Iic n) := fun i hi j hj hij => @strictMonoOn_Iic_of_lt_succ α βᵒᵈ _ _ ψ _ _ n hψ i hi j hj hij -theorem strictAnti_of_succ_lt [SuccOrder α] [IsSuccArchimedean α] - (hψ : ∀ m, ψ (succ m) < ψ m) : StrictAnti ψ := fun _ _ h ↦ - (strictAntiOn_Iic_of_succ_lt fun m _ ↦ hψ m) h.le le_rfl h - theorem strictMonoOn_Ici_of_pred_lt [PredOrder α] [IsPredArchimedean α] {n : α} (hψ : ∀ m, n < m → ψ (pred m) < ψ m) : StrictMonoOn ψ (Set.Ici n) := fun i hi j hj hij => @strictMonoOn_Iic_of_lt_succ αᵒᵈ βᵒᵈ _ _ ψ _ _ n hψ j hj i hi hij -theorem strictMono_of_pred_lt [PredOrder α] [IsPredArchimedean α] - (hψ : ∀ m, ψ (pred m) < ψ m) : StrictMono ψ := fun _ _ h ↦ - (strictMonoOn_Ici_of_pred_lt fun m _ ↦ hψ m) le_rfl h.le h - theorem strictAntiOn_Ici_of_lt_pred [PredOrder α] [IsPredArchimedean α] {n : α} (hψ : ∀ m, n < m → ψ m < ψ (pred m)) : StrictAntiOn ψ (Set.Ici n) := fun i hi j hj hij => @strictAntiOn_Iic_of_succ_lt αᵒᵈ βᵒᵈ _ _ ψ _ _ n hψ j hj i hi hij -theorem strictAnti_of_lt_pred [PredOrder α] [IsPredArchimedean α] - (hψ : ∀ m, ψ m < ψ (pred m)) : StrictAnti ψ := fun _ _ h ↦ - (strictAntiOn_Ici_of_lt_pred fun m _ ↦ hψ m) le_rfl h.le h - end SuccOrder section LinearOrder diff --git a/Mathlib/Order/SuccPred/Archimedean.lean b/Mathlib/Order/SuccPred/Archimedean.lean index 1518b8b1fdbca..4ba2d38f22bdc 100644 --- a/Mathlib/Order/SuccPred/Archimedean.lean +++ b/Mathlib/Order/SuccPred/Archimedean.lean @@ -396,3 +396,117 @@ instance Set.OrdConnected.isSuccArchimedean [SuccOrder α] [IsSuccArchimedean α inferInstanceAs (IsSuccArchimedean sᵒᵈᵒᵈ) end OrdConnected + +section Monotone +variable {α β : Type*} [PartialOrder α] [Preorder β] + +section SuccOrder +variable [SuccOrder α] [IsSuccArchimedean α] {s : Set α} {f : α → β} + +lemma monotoneOn_of_le_succ (hs : s.OrdConnected) + (hf : ∀ a, ¬ IsMax a → a ∈ s → succ a ∈ s → f a ≤ f (succ a)) : MonotoneOn f s := by + rintro a ha b hb hab + obtain ⟨n, rfl⟩ := exists_succ_iterate_of_le hab + clear hab + induction' n with n hn + · simp + rw [Function.iterate_succ_apply'] at hb ⊢ + have : succ^[n] a ∈ s := hs.1 ha hb ⟨le_succ_iterate .., le_succ _⟩ + by_cases hb' : IsMax (succ^[n] a) + · rw [succ_eq_iff_isMax.2 hb'] + exact hn this + · exact (hn this).trans (hf _ hb' this hb) + +lemma antitoneOn_of_succ_le (hs : s.OrdConnected) + (hf : ∀ a, ¬ IsMax a → a ∈ s → succ a ∈ s → f (succ a) ≤ f a) : AntitoneOn f s := + monotoneOn_of_le_succ (β := βᵒᵈ) hs hf + +lemma strictMonoOn_of_lt_succ (hs : s.OrdConnected) + (hf : ∀ a, ¬ IsMax a → a ∈ s → succ a ∈ s → f a < f (succ a)) : StrictMonoOn f s := by + rintro a ha b hb hab + obtain ⟨n, rfl⟩ := exists_succ_iterate_of_le hab.le + obtain _ | n := n + · simp at hab + apply not_isMax_of_lt at hab + induction' n with n hn + · simpa using hf _ hab ha hb + rw [Function.iterate_succ_apply'] at hb ⊢ + have : succ^[n + 1] a ∈ s := hs.1 ha hb ⟨le_succ_iterate .., le_succ _⟩ + by_cases hb' : IsMax (succ^[n + 1] a) + · rw [succ_eq_iff_isMax.2 hb'] + exact hn this + · exact (hn this).trans (hf _ hb' this hb) + +lemma strictAntiOn_of_succ_lt (hs : s.OrdConnected) + (hf : ∀ a, ¬ IsMax a → a ∈ s → succ a ∈ s → f (succ a) < f a) : StrictAntiOn f s := + strictMonoOn_of_lt_succ (β := βᵒᵈ) hs hf + +lemma monotone_of_le_succ (hf : ∀ a, ¬ IsMax a → f a ≤ f (succ a)) : Monotone f := by + simpa using monotoneOn_of_le_succ Set.ordConnected_univ (by simpa using hf) + +lemma antitone_of_succ_le (hf : ∀ a, ¬ IsMax a → f (succ a) ≤ f a) : Antitone f := by + simpa using antitoneOn_of_succ_le Set.ordConnected_univ (by simpa using hf) + +lemma strictMono_of_lt_succ (hf : ∀ a, ¬ IsMax a → f a < f (succ a)) : StrictMono f := by + simpa using strictMonoOn_of_lt_succ Set.ordConnected_univ (by simpa using hf) + +lemma strictAnti_of_succ_lt (hf : ∀ a, ¬ IsMax a → f (succ a) < f a) : StrictAnti f := by + simpa using strictAntiOn_of_succ_lt Set.ordConnected_univ (by simpa using hf) + +end SuccOrder + +section PredOrder +variable [PredOrder α] [IsPredArchimedean α] {s : Set α} {f : α → β} + +lemma monotoneOn_of_pred_le (hs : s.OrdConnected) + (hf : ∀ a, ¬ IsMin a → a ∈ s → pred a ∈ s → f (pred a) ≤ f a) : MonotoneOn f s := by + rintro a ha b hb hab + obtain ⟨n, rfl⟩ := exists_pred_iterate_of_le hab + clear hab + induction' n with n hn + · simp + rw [Function.iterate_succ_apply'] at ha ⊢ + have : pred^[n] b ∈ s := hs.1 ha hb ⟨pred_le _, pred_iterate_le ..⟩ + by_cases ha' : IsMin (pred^[n] b) + · rw [pred_eq_iff_isMin.2 ha'] + exact hn this + · exact (hn this).trans' (hf _ ha' this ha) + +lemma antitoneOn_of_le_pred (hs : s.OrdConnected) + (hf : ∀ a, ¬ IsMin a → a ∈ s → pred a ∈ s → f a ≤ f (pred a)) : AntitoneOn f s := + monotoneOn_of_pred_le (β := βᵒᵈ) hs hf + +lemma strictMonoOn_of_pred_lt (hs : s.OrdConnected) + (hf : ∀ a, ¬ IsMin a → a ∈ s → pred a ∈ s → f (pred a) < f a) : StrictMonoOn f s := by + rintro a ha b hb hab + obtain ⟨n, rfl⟩ := exists_pred_iterate_of_le hab.le + obtain _ | n := n + · simp at hab + apply not_isMin_of_lt at hab + induction' n with n hn + · simpa using hf _ hab hb ha + rw [Function.iterate_succ_apply'] at ha ⊢ + have : pred^[n + 1] b ∈ s := hs.1 ha hb ⟨pred_le _, pred_iterate_le ..⟩ + by_cases ha' : IsMin (pred^[n + 1] b) + · rw [pred_eq_iff_isMin.2 ha'] + exact hn this + · exact (hn this).trans' (hf _ ha' this ha) + +lemma strictAntiOn_of_lt_pred (hs : s.OrdConnected) + (hf : ∀ a, ¬ IsMin a → a ∈ s → pred a ∈ s → f a < f (pred a)) : StrictAntiOn f s := + strictMonoOn_of_pred_lt (β := βᵒᵈ) hs hf + +lemma monotone_of_pred_le (hf : ∀ a, ¬ IsMin a → f (pred a) ≤ f a) : Monotone f := by + simpa using monotoneOn_of_pred_le Set.ordConnected_univ (by simpa using hf) + +lemma antitone_of_le_pred (hf : ∀ a, ¬ IsMin a → f a ≤ f (pred a)) : Antitone f := by + simpa using antitoneOn_of_le_pred Set.ordConnected_univ (by simpa using hf) + +lemma strictMono_of_pred_lt (hf : ∀ a, ¬ IsMin a → f (pred a) < f a) : StrictMono f := by + simpa using strictMonoOn_of_pred_lt Set.ordConnected_univ (by simpa using hf) + +lemma strictAnti_of_lt_pred (hf : ∀ a, ¬ IsMin a → f a < f (pred a)) : StrictAnti f := by + simpa using strictAntiOn_of_lt_pred Set.ordConnected_univ (by simpa using hf) + +end PredOrder +end Monotone From 2afe56b3b843c4ade79d6af3903be2a47e9cdc16 Mon Sep 17 00:00:00 2001 From: Robin Carlier <57142648+robin-carlier@users.noreply.github.com> Date: Wed, 23 Oct 2024 20:36:10 +0000 Subject: [PATCH 345/505] chore: `prodAssoc` as an additive/multiplicative or linear equivalence (#17254) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Record multiplicativity/linearity of the equivalence `(E × F) × G ≃ E × (F × G)`. Adapt the construction of `LinearIsometryEquiv.prodAssoc` to use this. Co-authored-by: Robin Carlier --- Mathlib/Algebra/Group/Prod.lean | 19 +++++++++++++- .../Normed/Operator/LinearIsometry.lean | 12 +++------ Mathlib/LinearAlgebra/Prod.lean | 26 +++++++++++++++++++ 3 files changed, 48 insertions(+), 9 deletions(-) diff --git a/Mathlib/Algebra/Group/Prod.lean b/Mathlib/Algebra/Group/Prod.lean index 24f6bc4ccca86..02781d10d0067 100644 --- a/Mathlib/Algebra/Group/Prod.lean +++ b/Mathlib/Algebra/Group/Prod.lean @@ -587,7 +587,24 @@ theorem coe_prodComm : ⇑(prodComm : M × N ≃* N × M) = Prod.swap := theorem coe_prodComm_symm : ⇑(prodComm : M × N ≃* N × M).symm = Prod.swap := rfl -variable {M' N' : Type*} [MulOneClass M'] [MulOneClass N'] +variable [MulOneClass P] + +/-- The equivalence between `(M × N) × P` and `M × (N × P)` is multiplicative. -/ +@[to_additive prodAssoc + "The equivalence between `(M × N) × P` and `M × (N × P)` is additive."] +def prodAssoc : (M × N) × P ≃* M × (N × P) := + { Equiv.prodAssoc M N P with map_mul' := fun ⟨_, _⟩ ⟨_, _⟩ => rfl } + +@[to_additive (attr := simp) coe_prodAssoc] +theorem coe_prodAssoc : ⇑(prodAssoc : (M × N) × P ≃* M × (N × P)) = Equiv.prodAssoc M N P := + rfl + +@[to_additive (attr := simp) coe_prodAssoc_symm] +theorem coe_prodAssoc_symm : + ⇑(prodAssoc : (M × N) × P ≃* M × (N × P)).symm = (Equiv.prodAssoc M N P).symm := + rfl + +variable {M' : Type*} {N' : Type*} [MulOneClass N'] [MulOneClass M'] section diff --git a/Mathlib/Analysis/Normed/Operator/LinearIsometry.lean b/Mathlib/Analysis/Normed/Operator/LinearIsometry.lean index facb95753719a..efae1765145f5 100644 --- a/Mathlib/Analysis/Normed/Operator/LinearIsometry.lean +++ b/Mathlib/Analysis/Normed/Operator/LinearIsometry.lean @@ -945,16 +945,12 @@ variable (R E E₂ E₃) /-- The natural equivalence `(E × E₂) × E₃ ≃ E × (E₂ × E₃)` is a linear isometry. -/ def prodAssoc [Module R E₂] [Module R E₃] : (E × E₂) × E₃ ≃ₗᵢ[R] E × E₂ × E₃ := - { Equiv.prodAssoc E E₂ E₃ with - toFun := Equiv.prodAssoc E E₂ E₃ - invFun := (Equiv.prodAssoc E E₂ E₃).symm - map_add' := by simp [-_root_.map_add] -- Fix timeout from #8386 - map_smul' := by -- was `by simp` before #6057 caused that to time out. - rintro m ⟨⟨e, f⟩, g⟩ - simp only [Prod.smul_mk, Equiv.prodAssoc_apply, RingHom.id_apply] + { LinearEquiv.prodAssoc R E E₂ E₃ with norm_map' := by rintro ⟨⟨e, f⟩, g⟩ - simp only [LinearEquiv.coe_mk, Equiv.prodAssoc_apply, Prod.norm_def, max_assoc] } + simp only [LinearEquiv.prodAssoc_apply, AddEquiv.toEquiv_eq_coe, + Equiv.toFun_as_coe, EquivLike.coe_coe, AddEquiv.coe_prodAssoc, + Equiv.prodAssoc_apply, Prod.norm_def, max_assoc] } @[simp] theorem coe_prodAssoc [Module R E₂] [Module R E₃] : diff --git a/Mathlib/LinearAlgebra/Prod.lean b/Mathlib/LinearAlgebra/Prod.lean index c81e4d17cb82e..deb0c5569016f 100644 --- a/Mathlib/LinearAlgebra/Prod.lean +++ b/Mathlib/LinearAlgebra/Prod.lean @@ -657,6 +657,32 @@ theorem snd_comp_prodComm : end prodComm +/-- Product of modules is associative up to linear isomorphism. -/ +@[simps apply] +def prodAssoc (R M₁ M₂ M₃ : Type*) [Semiring R] + [AddCommMonoid M₁] [AddCommMonoid M₂] [AddCommMonoid M₃] + [Module R M₁] [Module R M₂] [Module R M₃] : ((M₁ × M₂) × M₃) ≃ₗ[R] (M₁ × (M₂ × M₃)) := + { AddEquiv.prodAssoc with + map_smul' := fun _r ⟨_m, _n⟩ => rfl } + +section prodAssoc + +variable {M₁ : Type*} +variable [Semiring R] [AddCommMonoid M₁] [AddCommMonoid M₂] [AddCommMonoid M₃] +variable [Module R M₁] [Module R M₂] [Module R M₃] + +theorem fst_comp_prodAssoc : + (LinearMap.fst R M₁ (M₂ × M₃)).comp (prodAssoc R M₁ M₂ M₃).toLinearMap = + (LinearMap.fst R M₁ M₂).comp (LinearMap.fst R (M₁ × M₂) M₃) := by + ext <;> simp + +theorem snd_comp_prodAssoc : + (LinearMap.snd R M₁ (M₂ × M₃)).comp (prodAssoc R M₁ M₂ M₃).toLinearMap = + (LinearMap.snd R M₁ M₂).prodMap (LinearMap.id : M₃ →ₗ[R] M₃):= by + ext <;> simp + +end prodAssoc + section variable (R M M₂ M₃ M₄) From 2d90b18f5f702e75641981f9ec2b5b8bdabc60d5 Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Wed, 23 Oct 2024 20:36:11 +0000 Subject: [PATCH 346/505] fix(Algebra/DirectSum/Internal): delete some instances (#18149) See [Zulip](https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/rewrite.20failure.20on.20Subalgebra/near/478555535) discussion. One of these instances was breaking user code. --- Mathlib/Algebra/DirectSum/Internal.lean | 6 ------ 1 file changed, 6 deletions(-) diff --git a/Mathlib/Algebra/DirectSum/Internal.lean b/Mathlib/Algebra/DirectSum/Internal.lean index 4c03bf25b2f72..99122f4b761d9 100644 --- a/Mathlib/Algebra/DirectSum/Internal.lean +++ b/Mathlib/Algebra/DirectSum/Internal.lean @@ -53,12 +53,6 @@ open DirectSum variable {ι : Type*} {σ S R : Type*} -instance AddCommMonoid.ofSubmonoidOnSemiring [Semiring R] [SetLike σ R] [AddSubmonoidClass σ R] - (A : ι → σ) : ∀ i, AddCommMonoid (A i) := fun i => by infer_instance - -instance AddCommGroup.ofSubgroupOnRing [Ring R] [SetLike σ R] [AddSubgroupClass σ R] (A : ι → σ) : - ∀ i, AddCommGroup (A i) := fun i => by infer_instance - theorem SetLike.algebraMap_mem_graded [Zero ι] [CommSemiring S] [Semiring R] [Algebra S R] (A : ι → Submodule S R) [SetLike.GradedOne A] (s : S) : algebraMap S R s ∈ A 0 := by rw [Algebra.algebraMap_eq_smul_one] From ca0f1392ff524a2a93b5a204c86a8dcdfb9d4fc7 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Wed, 23 Oct 2024 21:11:13 +0000 Subject: [PATCH 347/505] feat(SetTheory/Cardinal/Aleph): introduce `omega` function (#17669) We introduce the `preOmega` function enumerating the initial ordinals, as well as the more standard `omega` skipping over finite ordinals. --- Mathlib/SetTheory/Cardinal/Aleph.lean | 203 +++++++++++++++++---- Mathlib/SetTheory/Cardinal/Cofinality.lean | 14 +- Mathlib/SetTheory/Ordinal/Basic.lean | 3 - 3 files changed, 171 insertions(+), 49 deletions(-) diff --git a/Mathlib/SetTheory/Cardinal/Aleph.lean b/Mathlib/SetTheory/Cardinal/Aleph.lean index a7267f649efde..78df9ea23c7e0 100644 --- a/Mathlib/SetTheory/Cardinal/Aleph.lean +++ b/Mathlib/SetTheory/Cardinal/Aleph.lean @@ -1,30 +1,32 @@ /- Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Johannes Hölzl, Mario Carneiro, Floris van Doorn +Authors: Johannes Hölzl, Mario Carneiro, Floris van Doorn, Violeta Hernández Palacios -/ import Mathlib.Order.Bounded import Mathlib.SetTheory.Cardinal.PartENat -import Mathlib.SetTheory.Ordinal.Arithmetic +import Mathlib.SetTheory.Ordinal.Enum /-! -# Aleph and beth functions - -* The function `Cardinal.preAleph` gives the cardinals listed by their ordinal index. - `preAleph n = n`, `preAleph ω = ℵ₀`, `preAleph (ω + 1) = succ ℵ₀`, etc. - It is an order isomorphism between ordinals and cardinals. -* The function `Cardinal.aleph` gives the infinite cardinals listed by their - ordinal index. `aleph 0 = ℵ₀`, `aleph 1 = succ ℵ₀` is the first - uncountable cardinal, and so on. The notation `ω_` combines the latter with `Cardinal.ord`, - giving an enumeration of (infinite) initial ordinals. - Thus `ω_ 0 = ω` and `ω₁ = ω_ 1` is the first uncountable ordinal. +# Omega, aleph, and beth functions + +* The function `Ordinal.preOmega` enumerates the initial ordinals, i.e. the smallest ordinals with + any given cardinality. Thus `preOmega n = n`, `preOmega ω = ω`, `preOmega (ω + 1) = ω₁`, etc. + `Ordinal.omega` is the more standard function which skips over finite ordinals. +* The function `Cardinal.preAleph` is an order isomorphism between ordinals and cardinals. Thus + `preAleph n = n`, `preAleph ω = ℵ₀`, `preAleph (ω + 1) = ℵ₁`, etc. `Cardinal.aleph` is the more + standard function which skips over finite ordinals. * The function `Cardinal.beth` enumerates the Beth cardinals. `beth 0 = ℵ₀`, `beth (succ o) = 2 ^ beth o`, and for a limit ordinal `o`, `beth o` is the supremum of `beth a` for `a < o`. ## Notation -The following notation is scoped to the `Cardinal` namespace. +The following notations are scoped to the `Ordinal` namespace. + +- `ω_ o` is notation for `Ordinal.omega o`. `ω₁` is notation for `ω_ 1`. + +The following notations are scoped to the `Cardinal` namespace. - `ℵ_ o` is notation for `aleph o`. `ℵ₁` is notation for `ℵ_ 1`. - `ℶ_ o` is notation for `beth o`. The value `ℶ_ 1` equals the continuum `𝔠`, which is defined in @@ -91,8 +93,158 @@ def isInitialIso : {x // IsInitial x} ≃o Cardinal where right_inv x := card_ord x map_rel_iff' {a _} := a.2.card_le_card --- TODO: define `omega` as the enumerator function of `IsInitial`, and redefine --- `aleph x = (omega x).card`. +/-- The "pre-omega" function gives the initial ordinals listed by their ordinal index. +`preOmega n = n`, `preOmega ω = ω`, `preOmega (ω + 1) = ω₁`, etc. + +For the more common omega function skipping over finite ordinals, see `Ordinal.omega`. -/ +def preOmega : Ordinal.{u} ↪o Ordinal.{u} where + toFun := enumOrd {x | IsInitial x} + inj' _ _ h := enumOrd_injective not_bddAbove_isInitial h + map_rel_iff' := enumOrd_le_enumOrd not_bddAbove_isInitial + +theorem coe_preOmega : preOmega = enumOrd {x | IsInitial x} := + rfl + +theorem preOmega_strictMono : StrictMono preOmega := + preOmega.strictMono + +theorem preOmega_lt_preOmega {o₁ o₂ : Ordinal} : preOmega o₁ < preOmega o₂ ↔ o₁ < o₂ := + preOmega.lt_iff_lt + +theorem preOmega_le_preOmega {o₁ o₂ : Ordinal} : preOmega o₁ ≤ preOmega o₂ ↔ o₁ ≤ o₂ := + preOmega.le_iff_le + +theorem preOmega_max (o₁ o₂ : Ordinal) : preOmega (max o₁ o₂) = max (preOmega o₁) (preOmega o₂) := + preOmega.monotone.map_max + +theorem isInitial_preOmega (o : Ordinal) : IsInitial (preOmega o) := + enumOrd_mem not_bddAbove_isInitial o + +theorem le_preOmega_self (o : Ordinal) : o ≤ preOmega o := + preOmega_strictMono.le_apply + +@[simp] +theorem preOmega_zero : preOmega 0 = 0 := by + rw [coe_preOmega, enumOrd_zero] + exact csInf_eq_bot_of_bot_mem isInitial_zero + +@[simp] +theorem preOmega_natCast (n : ℕ) : preOmega n = n := by + induction n with + | zero => exact preOmega_zero + | succ n IH => + apply (le_preOmega_self _).antisymm' + apply enumOrd_succ_le not_bddAbove_isInitial (isInitial_natCast _) (IH.trans_lt _) + rw [Nat.cast_lt] + exact lt_succ n + +-- See note [no_index around OfNat.ofNat] +@[simp] +theorem preOmega_ofNat (n : ℕ) [n.AtLeastTwo] : preOmega (no_index (OfNat.ofNat n)) = n := + preOmega_natCast n + +theorem preOmega_le_of_forall_lt {o a : Ordinal} (ha : IsInitial a) (H : ∀ b < o, preOmega b < a) : + preOmega o ≤ a := + enumOrd_le_of_forall_lt ha H + +theorem isNormal_preOmega : IsNormal preOmega := by + rw [isNormal_iff_strictMono_limit] + refine ⟨preOmega_strictMono, fun o ho a ha ↦ + (preOmega_le_of_forall_lt (isInitial_ord _) fun b hb ↦ ?_).trans (ord_card_le a)⟩ + rw [← (isInitial_ord _).card_lt_card, card_ord] + apply lt_of_lt_of_le _ (card_le_card <| ha _ (ho.succ_lt hb)) + rw [(isInitial_preOmega _).card_lt_card, preOmega_lt_preOmega] + exact lt_succ b + +@[simp] +theorem range_preOmega : range preOmega = {x | IsInitial x} := + range_enumOrd not_bddAbove_isInitial + +theorem mem_range_preOmega_iff {x : Ordinal} : x ∈ range preOmega ↔ IsInitial x := by + rw [range_preOmega, mem_setOf] + +alias ⟨_, IsInitial.mem_range_preOmega⟩ := mem_range_preOmega_iff + +@[simp] +theorem preOmega_omega0 : preOmega ω = ω := by + simp_rw [← isNormal_preOmega.apply_omega0, preOmega_natCast, iSup_natCast] + +@[simp] +theorem omega0_le_preOmega_iff {x : Ordinal} : ω ≤ preOmega x ↔ ω ≤ x := by + conv_lhs => rw [← preOmega_omega0, preOmega_le_preOmega] + +@[simp] +theorem omega0_lt_preOmega_iff {x : Ordinal} : ω < preOmega x ↔ ω < x := by + conv_lhs => rw [← preOmega_omega0, preOmega_lt_preOmega] + +/-- The `omega` function gives the infinite initial ordinals listed by their ordinal index. +`omega 0 = ω`, `omega 1 = ω₁` is the first uncountable ordinal, and so on. + +This is not to be confused with the first infinite ordinal `Ordinal.omega0`. + +For a version including finite ordinals, see `Ordinal.preOmega`. -/ +def omega : Ordinal ↪o Ordinal := + (OrderEmbedding.addLeft ω).trans preOmega + +@[inherit_doc] +scoped notation "ω_ " => omega + +/-- `ω₁` is the first uncountable ordinal. -/ +scoped notation "ω₁" => ω_ 1 + +theorem omega_eq_preOmega (o : Ordinal) : ω_ o = preOmega (ω + o) := + rfl + +theorem omega_strictMono : StrictMono omega := + omega.strictMono + +theorem omega_lt_omega {o₁ o₂ : Ordinal} : ω_ o₁ < ω_ o₂ ↔ o₁ < o₂ := + omega.lt_iff_lt + +theorem omega_le_omega {o₁ o₂ : Ordinal} : ω_ o₁ ≤ ω_ o₂ ↔ o₁ ≤ o₂ := + omega.le_iff_le + +theorem omega_max (o₁ o₂ : Ordinal) : ω_ (max o₁ o₂) = max (ω_ o₁) (ω_ o₂) := + omega.monotone.map_max + +theorem isInitial_omega (o : Ordinal) : IsInitial (omega o) := + isInitial_preOmega _ + +theorem le_omega_self (o : Ordinal) : o ≤ omega o := + omega_strictMono.le_apply + +@[simp] +theorem omega_zero : ω_ 0 = ω := by + rw [omega_eq_preOmega, add_zero, preOmega_omega0] + +theorem omega0_le_omega (o : Ordinal) : ω ≤ ω_ o := by + rw [← omega_zero, omega_le_omega] + exact Ordinal.zero_le o + +theorem omega0_lt_omega1 : ω < ω₁ := by + rw [← omega_zero, omega_lt_omega] + exact zero_lt_one + +@[deprecated omega0_lt_omega1 (since := "2024-10-11")] +alias omega_lt_omega1 := omega0_lt_omega1 + +theorem isNormal_omega : IsNormal omega := + isNormal_preOmega.trans (isNormal_add_right _) + +@[simp] +theorem range_omega : range omega = {x | ω ≤ x ∧ IsInitial x} := by + ext x + constructor + · rintro ⟨a, rfl⟩ + exact ⟨omega0_le_omega a, isInitial_omega a⟩ + · rintro ⟨ha', ha⟩ + obtain ⟨a, rfl⟩ := ha.mem_range_preOmega + use a - ω + rw [omega0_le_preOmega_iff] at ha' + rw [omega_eq_preOmega, Ordinal.add_sub_cancel_of_le ha'] + +theorem mem_range_omega_iff {x : Ordinal} : x ∈ range omega ↔ ω ≤ x ∧ IsInitial x := by + rw [range_omega, mem_setOf] end Ordinal @@ -192,7 +344,7 @@ def aleph : Ordinal ↪o Cardinal := @[inherit_doc] scoped notation "ℵ_ " => aleph -/-- `ℵ₁` is the first uncountable ordinal. -/ +/-- `ℵ₁` is the first uncountable cardinal. -/ scoped notation "ℵ₁" => ℵ_ 1 theorem aleph_eq_preAleph (o : Ordinal) : ℵ_ o = preAleph (ω + o) := @@ -572,22 +724,3 @@ theorem beth_normal : IsNormal.{u} fun o => (beth o).ord := isNormal_beth end Cardinal - -/-! ### Omega ordinals -/ - -namespace Ordinal - -/-- -`ω_ o` is a notation for the *initial ordinal* of cardinality -`aleph o`. Thus, for example `ω_ 0 = ω`. --/ -scoped notation "ω_" o => ord <| aleph o - -/-- -`ω₁` is the first uncountable ordinal. --/ -scoped notation "ω₁" => ord <| aleph 1 - -lemma omega_lt_omega1 : ω < ω₁ := ord_aleph0.symm.trans_lt (ord_lt_ord.mpr (aleph0_lt_aleph_one)) - -end Ordinal diff --git a/Mathlib/SetTheory/Cardinal/Cofinality.lean b/Mathlib/SetTheory/Cardinal/Cofinality.lean index 9c31bb5780ec0..995e87d30575d 100644 --- a/Mathlib/SetTheory/Cardinal/Cofinality.lean +++ b/Mathlib/SetTheory/Cardinal/Cofinality.lean @@ -1198,22 +1198,14 @@ namespace Ordinal open Cardinal open scoped Ordinal --- TODO: generalize universes +-- TODO: generalize universes, and use ω₁. lemma iSup_sequence_lt_omega1 {α : Type u} [Countable α] - (o : α → Ordinal.{max u v}) (ho : ∀ n, o n < ω₁) : - iSup o < ω₁ := by + (o : α → Ordinal.{max u v}) (ho : ∀ n, o n < (aleph 1).ord) : + iSup o < (aleph 1).ord := by apply iSup_lt_ord_lift _ ho rw [Cardinal.isRegular_aleph_one.cof_eq] exact lt_of_le_of_lt mk_le_aleph0 aleph0_lt_aleph_one -set_option linter.deprecated false in -@[deprecated iSup_sequence_lt_omega1 (since := "2024-08-27")] -lemma sup_sequence_lt_omega1 {α} [Countable α] (o : α → Ordinal) (ho : ∀ n, o n < ω₁) : - sup o < ω₁ := by - apply sup_lt_ord_lift _ ho - rw [Cardinal.isRegular_aleph_one.cof_eq] - exact lt_of_le_of_lt mk_le_aleph0 aleph0_lt_aleph_one - end Ordinal end Omega1 diff --git a/Mathlib/SetTheory/Ordinal/Basic.lean b/Mathlib/SetTheory/Ordinal/Basic.lean index 9ccd58c41f693..f3bb4990f67f0 100644 --- a/Mathlib/SetTheory/Ordinal/Basic.lean +++ b/Mathlib/SetTheory/Ordinal/Basic.lean @@ -712,9 +712,6 @@ theorem lt_lift_iff {a : Ordinal.{u}} {b : Ordinal.{max u v}} : def omega0 : Ordinal.{u} := lift <| @type ℕ (· < ·) _ -@[deprecated Ordinal.omega0 (since := "2024-09-26")] -alias omega := omega0 - @[inherit_doc] scoped notation "ω" => Ordinal.omega0 From 5b9d3428db8ad17a2529d98bfc183fbe432e9798 Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Wed, 23 Oct 2024 21:11:15 +0000 Subject: [PATCH 348/505] feat(FieldTheory/Fixed): Finite group surjects onto automorphisms fixing the fixed field (#18143) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR shows that if a finite group `G` acts on a field `F`, then `G` surjects onto `F ≃ₐ[FixedPoints.subfield G F] F`. We already have this result in the case that `G` acts faithfully, so for the general case it's just a matter of quotienting by the kernel of the action. Co-authored-by: Thomas Browning --- Mathlib/Algebra/Algebra/Equiv.lean | 2 ++ Mathlib/FieldTheory/Fixed.lean | 44 ++++++++++++++++++++++++------ 2 files changed, 38 insertions(+), 8 deletions(-) diff --git a/Mathlib/Algebra/Algebra/Equiv.lean b/Mathlib/Algebra/Algebra/Equiv.lean index d1c80042e9fe6..f5e674a2f3566 100644 --- a/Mathlib/Algebra/Algebra/Equiv.lean +++ b/Mathlib/Algebra/Algebra/Equiv.lean @@ -801,6 +801,8 @@ theorem toAlgEquiv_injective [FaithfulSMul G A] : Function.Injective (MulSemiringAction.toAlgEquiv R A : G → A ≃ₐ[R] A) := fun _ _ h => eq_of_smul_eq_smul fun r => AlgEquiv.ext_iff.1 h r +variable (G) + /-- Each element of the group defines an algebra equivalence. This is a stronger version of `MulSemiringAction.toRingAut` and diff --git a/Mathlib/FieldTheory/Fixed.lean b/Mathlib/FieldTheory/Fixed.lean index 92f028d565270..dab929cb5fbef 100644 --- a/Mathlib/FieldTheory/Fixed.lean +++ b/Mathlib/FieldTheory/Fixed.lean @@ -17,7 +17,7 @@ This is the basis of the Fundamental Theorem of Galois Theory. Given a (finite) group `G` that acts on a field `F`, we define `FixedPoints.subfield G F`, the subfield consisting of elements of `F` fixed_points by every element of `G`. -This subfield is then normal and separable, and in addition (TODO) if `G` acts faithfully on `F` +This subfield is then normal and separable, and in addition if `G` acts faithfully on `F` then `finrank (FixedPoints.subfield G F) F = Fintype.card G`. ## Main Definitions @@ -311,8 +311,9 @@ theorem finrank_algHom (K : Type u) (V : Type v) [Field K] [Field V] [Algebra K namespace FixedPoints -theorem finrank_eq_card (G : Type u) (F : Type v) [Group G] [Field F] [Fintype G] - [MulSemiringAction G F] [FaithfulSMul G F] : +variable (G F : Type*) [Group G] [Field F] [MulSemiringAction G F] + +theorem finrank_eq_card [Fintype G] [FaithfulSMul G F] : finrank (FixedPoints.subfield G F) F = Fintype.card G := le_antisymm (FixedPoints.finrank_le_card G F) <| calc @@ -322,8 +323,7 @@ theorem finrank_eq_card (G : Type u) (F : Type v) [Group G] [Field F] [Fintype G _ = finrank (FixedPoints.subfield G F) F := finrank_linearMap_self _ _ _ /-- `MulSemiringAction.toAlgHom` is bijective. -/ -theorem toAlgHom_bijective (G : Type u) (F : Type v) [Group G] [Field F] [Finite G] - [MulSemiringAction G F] [FaithfulSMul G F] : +theorem toAlgHom_bijective [Finite G] [FaithfulSMul G F] : Function.Bijective (MulSemiringAction.toAlgHom _ _ : G → F →ₐ[subfield G F] F) := by cases nonempty_fintype G rw [Fintype.bijective_iff_injective_and_card] @@ -334,9 +334,37 @@ theorem toAlgHom_bijective (G : Type u) (F : Type v) [Group G] [Field F] [Finite · rw [← finrank_eq_card G F] exact LE.le.trans_eq (finrank_algHom _ F) (finrank_linearMap_self _ _ _) -/-- Bijection between G and algebra homomorphisms that fix the fixed points -/ -def toAlgHomEquiv (G : Type u) (F : Type v) [Group G] [Field F] [Finite G] [MulSemiringAction G F] - [FaithfulSMul G F] : G ≃ (F →ₐ[FixedPoints.subfield G F] F) := +/-- Bijection between `G` and algebra endomorphisms of `F` that fix the fixed points. -/ +def toAlgHomEquiv [Finite G] [FaithfulSMul G F] : G ≃ (F →ₐ[FixedPoints.subfield G F] F) := Equiv.ofBijective _ (toAlgHom_bijective G F) +/-- `MulSemiringAction.toAlgAut` is bijective. -/ +theorem toAlgAut_bijective [Finite G] [FaithfulSMul G F] : + Function.Bijective (MulSemiringAction.toAlgAut G (FixedPoints.subfield G F) F) := by + refine ⟨fun _ _ h ↦ (FixedPoints.toAlgHom_bijective G F).injective ?_, + fun f ↦ ((FixedPoints.toAlgHom_bijective G F).surjective f).imp (fun _ h ↦ ?_)⟩ <;> + rwa [DFunLike.ext_iff] at h ⊢ + +/-- Bijection between `G` and algebra automorphisms of `F` that fix the fixed points. -/ +def toAlgAutMulEquiv [Finite G] [FaithfulSMul G F] : G ≃* (F ≃ₐ[FixedPoints.subfield G F] F) := + MulEquiv.ofBijective _ (toAlgAut_bijective G F) + +/-- `MulSemiringAction.toAlgAut` is surjective. -/ +theorem toAlgAut_surjective [Finite G] : + Function.Surjective (MulSemiringAction.toAlgAut G (FixedPoints.subfield G F) F) := by + let f : G →* F ≃ₐ[FixedPoints.subfield G F] F := + MulSemiringAction.toAlgAut G (FixedPoints.subfield G F) F + let Q := G ⧸ f.ker + let _ : MulSemiringAction Q F := MulSemiringAction.compHom _ (QuotientGroup.kerLift f) + have : FaithfulSMul Q F := ⟨by + intro q₁ q₂ + refine Quotient.inductionOn₂' q₁ q₂ (fun g₁ g₂ h ↦ QuotientGroup.eq.mpr ?_) + rwa [MonoidHom.mem_ker, map_mul, map_inv, inv_mul_eq_one, AlgEquiv.ext_iff]⟩ + intro f + obtain ⟨q, hq⟩ := (toAlgAut_bijective Q F).surjective + (AlgEquiv.ofRingEquiv (f := f) (fun ⟨x, hx⟩ ↦ f.commutes' ⟨x, fun g ↦ hx g⟩)) + revert hq + refine QuotientGroup.induction_on q (fun g hg ↦ ⟨g, ?_⟩) + rwa [AlgEquiv.ext_iff] at hg ⊢ + end FixedPoints From b65e87ab079a01dd7909978d4a5d21a85457c252 Mon Sep 17 00:00:00 2001 From: Quang Dao Date: Wed, 23 Oct 2024 22:11:09 +0000 Subject: [PATCH 349/505] feat(Tactic/ReduceModChar): add fast modular exponentiation (#18030) This PR (done by Markus Himmel, with minor clean-up from Quang Dao) introduces fast modular exponentiation to the `reduce_mod_char` tactic. This allows for exponentiating large number reduced modulo `n` for some `n`. An example is: ``` example : (-126432 : ZMod 1235412223) ^ 12355342321 = 1001528716 := by reduce_mod_char ``` See the new tests for more examples. Co-authored-by: [Markus Himmel](https://github.com/TwoFX), [Quang Dao](https://github.com/quangvdao) Co-authored-by: Markus Himmel Co-authored-by: Matthew Ballard Co-authored-by: Quang Dao --- Mathlib.lean | 1 + Mathlib/NumberTheory/LucasPrimality.lean | 4 - Mathlib/Tactic.lean | 1 + Mathlib/Tactic/NormNum/Core.lean | 2 +- Mathlib/Tactic/NormNum/PowMod.lean | 123 +++++++++++++++++++++++ Mathlib/Tactic/ReduceModChar.lean | 57 +++++++++-- test/norm_num_ext.lean | 1 + test/reduce_mod_char.lean | 40 ++++++++ 8 files changed, 214 insertions(+), 15 deletions(-) create mode 100644 Mathlib/Tactic/NormNum/PowMod.lean diff --git a/Mathlib.lean b/Mathlib.lean index 8a9b819edad1f..c6b23b50974af 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4451,6 +4451,7 @@ import Mathlib.Tactic.NormNum.NatFib import Mathlib.Tactic.NormNum.NatSqrt import Mathlib.Tactic.NormNum.OfScientific import Mathlib.Tactic.NormNum.Pow +import Mathlib.Tactic.NormNum.PowMod import Mathlib.Tactic.NormNum.Prime import Mathlib.Tactic.NormNum.Result import Mathlib.Tactic.NthRewrite diff --git a/Mathlib/NumberTheory/LucasPrimality.lean b/Mathlib/NumberTheory/LucasPrimality.lean index 2f2f71e0b2bae..47887d769b36b 100644 --- a/Mathlib/NumberTheory/LucasPrimality.lean +++ b/Mathlib/NumberTheory/LucasPrimality.lean @@ -3,10 +3,6 @@ Copyright (c) 2020 Bolton Bailey. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Bolton Bailey -/ -import Mathlib.Data.Fintype.Basic -import Mathlib.GroupTheory.OrderOfElement -import Mathlib.Tactic.Zify -import Mathlib.Data.Nat.Totient import Mathlib.RingTheory.IntegralDomain /-! diff --git a/Mathlib/Tactic.lean b/Mathlib/Tactic.lean index fe71b6151623d..c72521006f0ab 100644 --- a/Mathlib/Tactic.lean +++ b/Mathlib/Tactic.lean @@ -180,6 +180,7 @@ import Mathlib.Tactic.NormNum.NatFib import Mathlib.Tactic.NormNum.NatSqrt import Mathlib.Tactic.NormNum.OfScientific import Mathlib.Tactic.NormNum.Pow +import Mathlib.Tactic.NormNum.PowMod import Mathlib.Tactic.NormNum.Prime import Mathlib.Tactic.NormNum.Result import Mathlib.Tactic.NthRewrite diff --git a/Mathlib/Tactic/NormNum/Core.lean b/Mathlib/Tactic/NormNum/Core.lean index 0c2aec57cf959..16ce0f37ed44d 100644 --- a/Mathlib/Tactic/NormNum/Core.lean +++ b/Mathlib/Tactic/NormNum/Core.lean @@ -94,7 +94,7 @@ def derive {α : Q(Type u)} (e : Q($α)) (post := false) : MetaM (Result e) := d trace[Tactic.norm_num] "{ext.name}:\n{e} ==> {new}" return new catch err => - trace[Tactic.norm_num] "{e} failed: {err.toMessageData}" + trace[Tactic.norm_num] "{ext.name} failed {e}: {err.toMessageData}" s.restore throwError "{e}: no norm_nums apply" diff --git a/Mathlib/Tactic/NormNum/PowMod.lean b/Mathlib/Tactic/NormNum/PowMod.lean new file mode 100644 index 0000000000000..f50964e19a2b3 --- /dev/null +++ b/Mathlib/Tactic/NormNum/PowMod.lean @@ -0,0 +1,123 @@ +/- +Copyright (c) 2023 Markus Himmel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Markus Himmel +-/ +import Mathlib.Tactic.NormNum.Pow + +/-! +# `norm_num` handling for expressions of the form `a ^ b % m`. + +These expressions can often be evaluated efficiently in cases where first evaluating `a ^ b` and +then reducing mod `m` is not feasible. We provide a function `evalNatPowMod` which is used by the +`reduce_mod_char` tactic to efficiently evaluate powers in rings with positive characteristic. + +The approach taken here is identical to (and copied from) the development in `NormNum/Pow.lean`. + +## TODO + +* Adapt the `norm_num` extensions for `Nat.mod` and `Int.emod` to efficiently evaluate expressions + of the form `a ^ b % m` using `evalNatPowMod`. + +-/ + +set_option autoImplicit true + +namespace Mathlib +open Lean hiding Rat mkRat +open Meta + +namespace Meta.NormNum +open Qq + +/-- Represents and proves equalities of the form `a^b % m = c` for natural numbers. -/ +structure IsNatPowModT (p : Prop) (a b m c : Nat) : Prop where + run' : p → Nat.mod (Nat.pow a b) m = c + +theorem IsNatPowModT.run + (p : IsNatPowModT (Nat.mod (Nat.pow a (nat_lit 1)) m = Nat.mod a m) a b m c) : + Nat.mod (Nat.pow a b) m = c := p.run' (congr_arg (fun x => x % m) (Nat.pow_one a)) + +theorem IsNatPowModT.trans (h1 : IsNatPowModT p a b m c) + (h2 : IsNatPowModT (Nat.mod (Nat.pow a b) m = c) a b' m c') : IsNatPowModT p a b' m c' := + ⟨h2.run' ∘ h1.run'⟩ + +theorem IsNatPowModT.bit0 : + IsNatPowModT (Nat.mod (Nat.pow a b) m = c) a (nat_lit 2 * b) m (Nat.mod (Nat.mul c c) m) := + ⟨fun h1 => by simp only [two_mul, Nat.pow_eq, pow_add, ← h1, Nat.mul_eq]; exact Nat.mul_mod ..⟩ + +theorem natPow_zero_natMod_zero : Nat.mod (Nat.pow a (nat_lit 0)) (nat_lit 0) = nat_lit 1 := by + with_unfolding_all rfl +theorem natPow_zero_natMod_one : Nat.mod (Nat.pow a (nat_lit 0)) (nat_lit 1) = nat_lit 0 := by + with_unfolding_all rfl +theorem natPow_zero_natMod_succ_succ : + Nat.mod (Nat.pow a (nat_lit 0)) (Nat.succ (Nat.succ m)) = nat_lit 1 := by + rw [natPow_zero] + apply Nat.mod_eq_of_lt + exact Nat.one_lt_succ_succ _ +theorem natPow_one_natMod : Nat.mod (Nat.pow a (nat_lit 1)) m = Nat.mod a m := by rw [natPow_one] + +theorem IsNatPowModT.bit1 : + IsNatPowModT (Nat.mod (Nat.pow a b) m = c) a (nat_lit 2 * b + 1) m + (Nat.mod (Nat.mul c (Nat.mod (Nat.mul c a) m)) m) := + ⟨by + rintro rfl + show a ^ (2 * b + 1) % m = (a ^ b % m) * ((a ^ b % m * a) % m) % m + rw [pow_add, two_mul, pow_add, pow_one, Nat.mul_mod (a ^ b % m) a, Nat.mod_mod, + ← Nat.mul_mod (a ^ b) a, ← Nat.mul_mod, mul_assoc]⟩ + +/-- Evaluates and proves `a^b % m` for natural numbers using fast modular exponentiation. -/ +partial def evalNatPowMod (a b m : Q(ℕ)) : (c : Q(ℕ)) × Q(Nat.mod (Nat.pow $a $b) $m = $c) := + if b.natLit! = 0 then + haveI : $b =Q 0 := ⟨⟩ + if m.natLit! = 0 then -- a ^ 0 % 0 = 1 + haveI : $m =Q 0 := ⟨⟩ + ⟨q(nat_lit 1), q(natPow_zero_natMod_zero)⟩ + else + have m' : Q(ℕ) := mkRawNatLit (m.natLit! - 1) + if m'.natLit! = 0 then -- a ^ 0 % 1 = 0 + haveI : $m =Q 1 := ⟨⟩ + ⟨q(nat_lit 0), q(natPow_zero_natMod_one)⟩ + else -- a ^ 0 % m = 1 + have m'' : Q(ℕ) := mkRawNatLit (m'.natLit! - 1) + haveI : $m =Q Nat.succ (Nat.succ $m'') := ⟨⟩ + ⟨q(nat_lit 1), q(natPow_zero_natMod_succ_succ)⟩ + else if b.natLit! = 1 then -- a ^ 1 % m = a % m + have c : Q(ℕ) := mkRawNatLit (a.natLit! % m.natLit!) + haveI : $b =Q 1 := ⟨⟩ + haveI : $c =Q Nat.mod $a $m := ⟨⟩ + ⟨c, q(natPow_one_natMod)⟩ + else + have c₀ : Q(ℕ) := mkRawNatLit (a.natLit! % m.natLit!) + haveI : $c₀ =Q Nat.mod $a $m := ⟨⟩ + let ⟨c, p⟩ := go b.natLit!.log2 a m (mkRawNatLit 1) c₀ b _ .rfl + ⟨c, q(($p).run)⟩ +where + /-- Invariants: `a ^ b₀ % m = c₀`, `depth > 0`, `b >>> depth = b₀` -/ + go (depth : Nat) (a m b₀ c₀ b : Q(ℕ)) + (p : Q(Prop)) (hp : $p =Q (Nat.mod (Nat.pow $a $b₀) $m = $c₀)) : + (c : Q(ℕ)) × Q(IsNatPowModT $p $a $b $m $c) := + let b' := b.natLit! + let m' := m.natLit! + if depth ≤ 1 then + let a' := a.natLit! + let c₀' := c₀.natLit! + if b' &&& 1 == 0 then + have c : Q(ℕ) := mkRawNatLit ((c₀' * c₀') % m') + haveI : $c =Q Nat.mod (Nat.mul $c₀ $c₀) $m := ⟨⟩ + haveI : $b =Q 2 * $b₀ := ⟨⟩ + ⟨c, q(IsNatPowModT.bit0)⟩ + else + have c : Q(ℕ) := mkRawNatLit ((c₀' * ((c₀' * a') % m')) % m') + haveI : $c =Q Nat.mod (Nat.mul $c₀ (Nat.mod (Nat.mul $c₀ $a) $m)) $m := ⟨⟩ + haveI : $b =Q 2 * $b₀ + 1 := ⟨⟩ + ⟨c, q(IsNatPowModT.bit1)⟩ + else + let d := depth >>> 1 + have hi : Q(ℕ) := mkRawNatLit (b' >>> d) + let ⟨c1, p1⟩ := go (depth - d) a m b₀ c₀ hi p (by exact hp) + let ⟨c2, p2⟩ := go d a m hi c1 b q(Nat.mod (Nat.pow $a $hi) $m = $c1) ⟨⟩ + ⟨c2, q(($p1).trans $p2)⟩ +end NormNum +end Meta +end Mathlib diff --git a/Mathlib/Tactic/ReduceModChar.lean b/Mathlib/Tactic/ReduceModChar.lean index 8a3886895ef48..859630e1b3a4f 100644 --- a/Mathlib/Tactic/ReduceModChar.lean +++ b/Mathlib/Tactic/ReduceModChar.lean @@ -6,6 +6,7 @@ Authors: Anne Baanen import Mathlib.Data.ZMod.Basic import Mathlib.RingTheory.Polynomial.Basic import Mathlib.Tactic.NormNum.DivMod +import Mathlib.Tactic.NormNum.PowMod import Mathlib.Tactic.ReduceModChar.Ext /-! @@ -44,26 +45,62 @@ open Mathlib.Meta.NormNum variable {u : Level} +lemma CharP.isInt_of_mod {e' r : ℤ} {α : Type*} [Ring α] {n n' : ℕ} (inst : CharP α n) {e : α} + (he : IsInt e e') (hn : IsNat n n') (h₂ : IsInt (e' % n') r) : IsInt e r := + ⟨by rw [he.out, CharP.intCast_eq_intCast_mod α n, show n = n' from hn.out, h₂.out, Int.cast_id]⟩ + +lemma CharP.isNat_pow {α} [Semiring α] : ∀ {f : α → ℕ → α} {a : α} {a' b b' c n n' : ℕ}, + CharP α n → f = HPow.hPow → IsNat a a' → IsNat b b' → IsNat n n' → + Nat.mod (Nat.pow a' b') n' = c → IsNat (f a b) c + | _, _, a, _, b, _, _, n, _, rfl, ⟨h⟩, ⟨rfl⟩, ⟨rfl⟩, rfl => ⟨by + rw [h, Nat.cast_id, Nat.pow_eq, ← Nat.cast_pow, CharP.natCast_eq_natCast_mod α n] + rfl⟩ + +/-- Evaluates `e` to an integer using `norm_num` and reduces the result modulo `n`. -/ +def normBareNumeral {α : Q(Type u)} (n n' : Q(ℕ)) (pn : Q(IsNat «$n» «$n'»)) + (e : Q($α)) (_ : Q(Ring $α)) (instCharP : Q(CharP $α $n)) : MetaM (Result e) := do + let ⟨ze, ne, pe⟩ ← Result.toInt _ (← Mathlib.Meta.NormNum.derive e) + let rr ← evalIntMod.go _ _ ze q(IsInt.raw_refl $ne) _ <| + .isNat q(instAddMonoidWithOne) _ q(isNat_natCast _ _ (IsNat.raw_refl $n')) + let ⟨zr, nr, pr⟩ ← rr.toInt _ + return .isInt _ nr zr q(CharP.isInt_of_mod $instCharP $pe $pn $pr) + +mutual + + /-- Given an expression of the form `a ^ b` in a ring of characteristic `n`, reduces `a` + modulo `n` recursively and then calculates `a ^ b` using fast modular exponentiation. -/ + partial def normPow {α : Q(Type u)} (n n' : Q(ℕ)) (pn : Q(IsNat «$n» «$n'»)) (e : Q($α)) + (_ : Q(Ring $α)) (instCharP : Q(CharP $α $n)) : MetaM (Result e) := do + let .app (.app (f : Q($α → ℕ → $α)) (a : Q($α))) (b : Q(ℕ)) ← whnfR e | failure + let .isNat sα na pa ← normIntNumeral' n n' pn a _ instCharP | failure + let ⟨nb, pb⟩ ← Mathlib.Meta.NormNum.deriveNat b q(instAddMonoidWithOneNat) + guard <|← withNewMCtxDepth <| isDefEq f q(HPow.hPow (α := $α)) + haveI' : $e =Q $a ^ $b := ⟨⟩ + haveI' : $f =Q HPow.hPow := ⟨⟩ + have ⟨c, r⟩ := evalNatPowMod na nb n' + assumeInstancesCommute + return .isNat sα c q(CharP.isNat_pow (f := $f) $instCharP (.refl $f) $pa $pb $pn $r) + + /-- If `e` is of the form `a ^ b`, reduce it using fast modular exponentation, otherwise + reduce it using `norm_num`. -/ + partial def normIntNumeral' {α : Q(Type u)} (n n' : Q(ℕ)) (pn : Q(IsNat «$n» «$n'»)) + (e : Q($α)) (_ : Q(Ring $α)) (instCharP : Q(CharP $α $n)) : MetaM (Result e) := + normPow n n' pn e _ instCharP <|> normBareNumeral n n' pn e _ instCharP + +end + lemma CharP.intCast_eq_mod (R : Type _) [Ring R] (p : ℕ) [CharP R p] (k : ℤ) : (k : R) = (k % p : ℤ) := by calc (k : R) = ↑(k % p + p * (k / p)) := by rw [Int.emod_add_ediv] _ = ↑(k % p) := by simp [CharP.cast_eq_zero R] -lemma CharP.isInt_of_mod {e' r : ℤ} {α : Type _} [Ring α] {n n' : ℕ} (inst : CharP α n) {e : α} - (he : IsInt e e') (hn : IsNat n n') (h₂ : IsInt (e' % n') r) : IsInt e r := - ⟨by rw [he.out, CharP.intCast_eq_mod α n, show n = n' from hn.out, h₂.out, Int.cast_id]⟩ - /-- Given an integral expression `e : t` such that `t` is a ring of characteristic `n`, reduce `e` modulo `n`. -/ -partial def normIntNumeral {α : Q(Type u)} (n : Q(ℕ)) (e : Q($α)) (instRing : Q(Ring $α)) +partial def normIntNumeral {α : Q(Type u)} (n : Q(ℕ)) (e : Q($α)) (_ : Q(Ring $α)) (instCharP : Q(CharP $α $n)) : MetaM (Result e) := do - let ⟨ze, ne, pe⟩ ← Result.toInt instRing (← Mathlib.Meta.NormNum.derive e) let ⟨n', pn⟩ ← deriveNat n q(instAddMonoidWithOneNat) - let rr ← evalIntMod.go _ _ ze q(IsInt.raw_refl $ne) _ <| - .isNat q(instAddMonoidWithOne) _ q(isNat_natCast _ _ (IsNat.raw_refl $n')) - let ⟨zr, nr, pr⟩ ← rr.toInt q(Int.instRing) - return .isInt instRing nr zr q(CharP.isInt_of_mod $instCharP $pe $pn $pr) + normIntNumeral' n n' pn e _ instCharP lemma CharP.neg_eq_sub_one_mul {α : Type _} [Ring α] (n : ℕ) (inst : CharP α n) (b : α) (a : ℕ) (a' : α) (p : IsNat (n - 1 : α) a) (pa : a = a') : diff --git a/test/norm_num_ext.lean b/test/norm_num_ext.lean index 4db23b8b748f9..d8ac0c6ac34e3 100644 --- a/test/norm_num_ext.lean +++ b/test/norm_num_ext.lean @@ -11,6 +11,7 @@ import Mathlib.Tactic.NormNum.NatFib import Mathlib.Tactic.NormNum.NatSqrt import Mathlib.Tactic.NormNum.Prime import Mathlib.Tactic.NormNum.LegendreSymbol +import Mathlib.Tactic.NormNum.Pow /-! # Tests for `norm_num` extensions diff --git a/test/reduce_mod_char.lean b/test/reduce_mod_char.lean index 597a56853552e..4180cc0fd347e 100644 --- a/test/reduce_mod_char.lean +++ b/test/reduce_mod_char.lean @@ -29,6 +29,46 @@ example : (X ^ 2 - 3 : (ZMod 4)[X]) = X ^ 2 + 1 := by reduce_mod_char -- Cleaning up `1 * X` and `X + 0`: example : (5 * X ^ 2 - 3 * X + 4 : (ZMod 4)[X]) = X ^ 2 + X := by reduce_mod_char +-- Exponentiation: +example : (11 : ZMod 987654319) ^ 987654318 = 1 := by reduce_mod_char +example : (-126432 : ZMod 1235412223) ^ 12355342321 = 1001528716 := by reduce_mod_char +example : (((((5 : ZMod 1235412223) ^ 5) ^ 5) ^ 5) ^ 5) ^ 5 = 806432269 := by reduce_mod_char +example : (2 : ZMod 75) ^ 0 = 1 := by reduce_mod_char +example : (0 : ZMod 34523432) ^ 0 = 1 := by reduce_mod_char +example : (0 : ZMod 56) ^ 90000000000000000 = 0 := by reduce_mod_char +example : (1 : ZMod 119) ^ 433440000000000000000000 = 1 := by reduce_mod_char +example : (75 : ZMod 111) ^ 1 = 75 := by reduce_mod_char +example : (23 : ZMod 19) ^ 1 = 4 := by reduce_mod_char +example : (1923423451 : ZMod 2) ^ 81000000000000000000 = 1 := by reduce_mod_char +example : (19 : ZMod 1) ^ 810000000000000000000 = 0 := by reduce_mod_char +example : (23423432049230423 : ZMod 0) ^ 0 = 1 := by reduce_mod_char +example : (23423432049230423 : ZMod 1) ^ 0 = 0 := by reduce_mod_char + +-- A 1024-bit prime number +set_option linter.style.longLine false in +abbrev p : Nat := 0xb10b8f96a080e01dde92de5eae5d54ec52c99fbcfb06a3c69a6a9dca52d23b616073e28675a23d189838ef1e2ee652c013ecb4aea906112324975c3cd49b83bfaccbdd7d90c4bd7098488e9c219a73724effd6fae5644738faa31a4ff55bccc0a151af5f0dc8b4bd45bf37df365c1a65e68cfda76d4da708df1fb2bc2e4a4371 + +set_option linter.style.longLine false in +abbrev g : Nat := 0xa4d1cbd5c3fd34126765a442efb99905f8104dd258ac507fd6406cff14266d31266fea1e5c41564b777e690f5504f213160217b4b01b886a5e91547f9e2749f4d7fbd7d3b9a92ee1909d0d2263f80a76a6a24c087a091f531dbf0a0169b6a28ad662a4d18e73afa32d779d5918d08bc8858f4dcef97c2a24855e6eeb22b3b2e5 + +example : (g : ZMod p) ^ (p - 1) = 1 := by reduce_mod_char + +-- The 20th Mersenne prime +abbrev q : Nat := 2 ^ 4423 - 1 + +-- This takes a little time but the 21st Mersenne prime is too large +set_option exponentiation.threshold 10000 in +example : (3 : ZMod q) ^ (q - 1) = 1 := by reduce_mod_char + +-- We don't unfold semi-reducible definitions +def q' := 2 +/-- +error: unsolved goals +⊢ 3 ^ (q' - 1) = 1 +-/ +#guard_msgs in +example : (3 : ZMod q') ^ (q' - 1) = 1 := by reduce_mod_char + -- Rewriting hypotheses: example (a : ZMod 7) (h : a + 7 = 2) : a = 2 := by reduce_mod_char at h From 82048966921ac92f86b5f91ff9991d655e6a6691 Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Wed, 23 Oct 2024 22:11:10 +0000 Subject: [PATCH 350/505] feat(RingTheory/Ideal/Pointwise): Add `IsPrime.smul` and `IsPrime.smul_iff` (#18147) This PR adds lemmas `IsPrime.smul` and `IsPrime.smul_iff`. --- Mathlib/RingTheory/Ideal/Maps.lean | 3 ++- Mathlib/RingTheory/Ideal/Pointwise.lean | 8 ++++++++ 2 files changed, 10 insertions(+), 1 deletion(-) diff --git a/Mathlib/RingTheory/Ideal/Maps.lean b/Mathlib/RingTheory/Ideal/Maps.lean index 84caddfcb5d16..0462faa6fbffb 100644 --- a/Mathlib/RingTheory/Ideal/Maps.lean +++ b/Mathlib/RingTheory/Ideal/Maps.lean @@ -212,8 +212,9 @@ theorem comap_sInf (s : Set (Ideal S)) : (sInf s).comap f = ⨅ I ∈ s, (I : Id theorem comap_sInf' (s : Set (Ideal S)) : (sInf s).comap f = ⨅ I ∈ comap f '' s, I := _root_.trans (comap_sInf f s) (by rw [iInf_image]) +/-- Variant of `Ideal.IsPrime.comap` where ideal is explicit rather than implicit. -/ theorem comap_isPrime [H : IsPrime K] : IsPrime (comap f K) := - ⟨comap_ne_top f H.ne_top, fun {x y} h => H.mem_or_mem <| by rwa [mem_comap, map_mul] at h⟩ + H.comap f variable {I J K L} diff --git a/Mathlib/RingTheory/Ideal/Pointwise.lean b/Mathlib/RingTheory/Ideal/Pointwise.lean index dad3393be0088..cde9b62eabc90 100644 --- a/Mathlib/RingTheory/Ideal/Pointwise.lean +++ b/Mathlib/RingTheory/Ideal/Pointwise.lean @@ -127,6 +127,14 @@ theorem pointwise_smul_subset_iff {a : M} {S T : Ideal R} : a • S ≤ T ↔ S theorem subset_pointwise_smul_iff {a : M} {S T : Ideal R} : S ≤ a • T ↔ a⁻¹ • S ≤ T := by rw [← pointwise_smul_le_pointwise_smul_iff (a := a⁻¹), inv_smul_smul] +theorem IsPrime.smul {I : Ideal R} [H : I.IsPrime] (g : M) : (g • I).IsPrime := by + rw [I.pointwise_smul_eq_comap] + apply H.comap + +@[simp] +theorem IsPrime.smul_iff {I : Ideal R} (g : M) : (g • I).IsPrime ↔ I.IsPrime := + ⟨fun H ↦ inv_smul_smul g I ▸ H.smul g⁻¹, fun H ↦ H.smul g⟩ + /-! TODO: add `equivSMul` like we have for subgroup. -/ end Group From f375907aff43996d648bf90dc95bc2c89e50626e Mon Sep 17 00:00:00 2001 From: Xavier Roblot Date: Wed, 23 Oct 2024 22:40:40 +0000 Subject: [PATCH 351/505] feat(Analysis/Complex): Prove that `Complex.ofReal` is a uniform embedding (#17943) We deduce from that the other direction of `Filter.Tendsto.ofReal`. This PR is part of the proof of the Analytic Class Number Formula. --- Mathlib/Analysis/Complex/Basic.lean | 9 ++++++++- 1 file changed, 8 insertions(+), 1 deletion(-) diff --git a/Mathlib/Analysis/Complex/Basic.lean b/Mathlib/Analysis/Complex/Basic.lean index 55726f263e5bf..eefdc2d5f0b00 100644 --- a/Mathlib/Analysis/Complex/Basic.lean +++ b/Mathlib/Analysis/Complex/Basic.lean @@ -358,9 +358,16 @@ theorem isometry_ofReal : Isometry ((↑) : ℝ → ℂ) := theorem continuous_ofReal : Continuous ((↑) : ℝ → ℂ) := ofRealLI.continuous +theorem isUniformEmbedding_ofReal : IsUniformEmbedding ((↑) : ℝ → ℂ) := + ofRealLI.isometry.isUniformEmbedding + +theorem _root_.Filter.tendsto_ofReal_iff {α : Type*} {l : Filter α} {f : α → ℝ} {x : ℝ} : + Tendsto (fun x ↦ (f x : ℂ)) l (𝓝 (x : ℂ)) ↔ Tendsto f l (𝓝 x) := + isUniformEmbedding_ofReal.toIsClosedEmbedding.tendsto_nhds_iff.symm + lemma _root_.Filter.Tendsto.ofReal {α : Type*} {l : Filter α} {f : α → ℝ} {x : ℝ} (hf : Tendsto f l (𝓝 x)) : Tendsto (fun x ↦ (f x : ℂ)) l (𝓝 (x : ℂ)) := - (continuous_ofReal.tendsto _).comp hf + tendsto_ofReal_iff.mpr hf /-- The only continuous ring homomorphism from `ℝ` to `ℂ` is the identity. -/ theorem ringHom_eq_ofReal_of_continuous {f : ℝ →+* ℂ} (h : Continuous f) : f = ofRealHom := by From 45deead3df330bee30208e5e46f8f655df30c4ea Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Bj=C3=B8rn=20Kjos-Hanssen?= Date: Wed, 23 Oct 2024 23:43:57 +0000 Subject: [PATCH 352/505] feat: add First Derivative Test from calculus (#16802) Add proofs of the First Derivative Test, in max and min forms, using neighborhood filters. New branch due to merge conflict in the [old](https://github.com/leanprover-community/mathlib4/tree/bjoernkjoshanssen_firstderivativetest) branch. --- Mathlib.lean | 2 + .../Calculus/FirstDerivativeTest.lean | 114 ++++++++++++++++++ Mathlib/Topology/ContinuousOn.lean | 39 ++++++ Mathlib/Topology/Order/LocalExtr.lean | 18 +++ Mathlib/Topology/Order/OrderClosedExtr.lean | 49 ++++++++ 5 files changed, 222 insertions(+) create mode 100644 Mathlib/Analysis/Calculus/FirstDerivativeTest.lean create mode 100644 Mathlib/Topology/Order/OrderClosedExtr.lean diff --git a/Mathlib.lean b/Mathlib.lean index c6b23b50974af..e9a3b499c9093 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1056,6 +1056,7 @@ import Mathlib.Analysis.Calculus.FDeriv.RestrictScalars import Mathlib.Analysis.Calculus.FDeriv.Star import Mathlib.Analysis.Calculus.FDeriv.Symmetric import Mathlib.Analysis.Calculus.FDeriv.WithLp +import Mathlib.Analysis.Calculus.FirstDerivativeTest import Mathlib.Analysis.Calculus.FormalMultilinearSeries import Mathlib.Analysis.Calculus.Gradient.Basic import Mathlib.Analysis.Calculus.Implicit @@ -4882,6 +4883,7 @@ import Mathlib.Topology.Order.MonotoneContinuity import Mathlib.Topology.Order.MonotoneConvergence import Mathlib.Topology.Order.NhdsSet import Mathlib.Topology.Order.OrderClosed +import Mathlib.Topology.Order.OrderClosedExtr import Mathlib.Topology.Order.PartialSups import Mathlib.Topology.Order.Priestley import Mathlib.Topology.Order.ProjIcc diff --git a/Mathlib/Analysis/Calculus/FirstDerivativeTest.lean b/Mathlib/Analysis/Calculus/FirstDerivativeTest.lean new file mode 100644 index 0000000000000..381835d729b57 --- /dev/null +++ b/Mathlib/Analysis/Calculus/FirstDerivativeTest.lean @@ -0,0 +1,114 @@ +/- +Copyright (c) 2024 Bjørn Kjos-Hanssen. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Bjørn Kjos-Hanssen, Patrick Massot, Floris van Doorn +-/ +import Mathlib.Analysis.Calculus.MeanValue +import Mathlib.Topology.Order.OrderClosedExtr +/-! +# The First-Derivative Test + +We prove the first-derivative test in the strong form given on [Wikipedia](https://en.wikipedia.org/wiki/Derivative_test#First-derivative_test). + +The test is proved over the real numbers ℝ +using `monotoneOn_of_deriv_nonneg` from [Mathlib.Analysis.Calculus.MeanValue]. + +## Main results + +* `isLocalMax_of_deriv_Ioo`: Suppose `f` is a real-valued function of a real variable + defined on some interval containing the point `a`. + Further suppose that `f` is continuous at `a` and differentiable on some open interval + containing `a`, except possibly at `a` itself. + + If there exists a positive number `r > 0` such that for every `x` in `(a − r, a)` + we have `f′(x) ≥ 0`, and for every `x` in `(a, a + r)` we have `f′(x) ≤ 0`, + then `f` has a local maximum at `a`. + +* `isLocalMin_of_deriv_Ioo`: The dual of `first_derivative_max`, for minima. + +* `isLocalMax_of_deriv`: 1st derivative test for maxima using filters. + +* `isLocalMin_of_deriv`: 1st derivative test for minima using filters. + +## Tags + +derivative test, calculus +-/ + +open Set Topology + + + /-- The First-Derivative Test from calculus, maxima version. + Suppose `a < b < c`, `f : ℝ → ℝ` is continuous at `b`, + the derivative `f'` is nonnegative on `(a,b)`, and + the derivative `f'` is nonpositive on `(b,c)`. Then `f` has a local maximum at `a`. -/ +lemma isLocalMax_of_deriv_Ioo {f : ℝ → ℝ} {a b c : ℝ} (g₀ : a < b) (g₁ : b < c) + (h : ContinuousAt f b) + (hd₀ : DifferentiableOn ℝ f (Ioo a b)) + (hd₁ : DifferentiableOn ℝ f (Ioo b c)) + (h₀ : ∀ x ∈ Ioo a b, 0 ≤ deriv f x) + (h₁ : ∀ x ∈ Ioo b c, deriv f x ≤ 0) : IsLocalMax f b := + have hIoc : ContinuousOn f (Ioc a b) := + Ioo_union_right g₀ ▸ hd₀.continuousOn.union_continuousAt (isOpen_Ioo (a := a) (b := b)) + (by simp_all) + have hIco : ContinuousOn f (Ico b c) := + Ioo_union_left g₁ ▸ hd₁.continuousOn.union_continuousAt (isOpen_Ioo (a := b) (b := c)) + (by simp_all) + isLocalMax_of_mono_anti g₀ g₁ + (monotoneOn_of_deriv_nonneg (convex_Ioc a b) hIoc (by simp_all) (by simp_all)) + (antitoneOn_of_deriv_nonpos (convex_Ico b c) hIco (by simp_all) (by simp_all)) + + +/-- The First-Derivative Test from calculus, minima version. -/ +lemma isLocalMin_of_deriv_Ioo {f : ℝ → ℝ} {a b c : ℝ} + (g₀ : a < b) (g₁ : b < c) (h : ContinuousAt f b) + (hd₀ : DifferentiableOn ℝ f (Ioo a b)) (hd₁ : DifferentiableOn ℝ f (Ioo b c)) + (h₀ : ∀ x ∈ Ioo a b, deriv f x ≤ 0) + (h₁ : ∀ x ∈ Ioo b c, 0 ≤ deriv f x) : IsLocalMin f b := by + have := isLocalMax_of_deriv_Ioo (f := -f) g₀ g₁ + (by simp_all) hd₀.neg hd₁.neg + (fun x hx => deriv.neg (f := f) ▸ Left.nonneg_neg_iff.mpr <|h₀ x hx) + (fun x hx => deriv.neg (f := f) ▸ Left.neg_nonpos_iff.mpr <|h₁ x hx) + exact (neg_neg f) ▸ IsLocalMax.neg this + + /-- The First-Derivative Test from calculus, maxima version, + expressed in terms of left and right filters. -/ +lemma isLocalMax_of_deriv' {f : ℝ → ℝ} {b : ℝ} (h : ContinuousAt f b) + (hd₀ : ∀ᶠ x in 𝓝[<] b, DifferentiableAt ℝ f x) (hd₁ : ∀ᶠ x in 𝓝[>] b, DifferentiableAt ℝ f x) + (h₀ : ∀ᶠ x in 𝓝[<] b, 0 ≤ deriv f x) (h₁ : ∀ᶠ x in 𝓝[>] b, deriv f x ≤ 0) : + IsLocalMax f b := by + obtain ⟨a,ha⟩ := (nhdsWithin_Iio_basis' ⟨b - 1, sub_one_lt b⟩).eventually_iff.mp <| hd₀.and h₀ + obtain ⟨c,hc⟩ := (nhdsWithin_Ioi_basis' ⟨b + 1, lt_add_one b⟩).eventually_iff.mp <| hd₁.and h₁ + exact isLocalMax_of_deriv_Ioo ha.1 hc.1 h + (fun _ hx => (ha.2 hx).1.differentiableWithinAt) + (fun _ hx => (hc.2 hx).1.differentiableWithinAt) + (fun _ hx => (ha.2 hx).2) (fun x hx => (hc.2 hx).2) + + /-- The First-Derivative Test from calculus, minima version, + expressed in terms of left and right filters. -/ +lemma isLocalMin_of_deriv' {f : ℝ → ℝ} {b : ℝ} (h : ContinuousAt f b) + (hd₀ : ∀ᶠ x in 𝓝[<] b, DifferentiableAt ℝ f x) (hd₁ : ∀ᶠ x in 𝓝[>] b, DifferentiableAt ℝ f x) + (h₀ : ∀ᶠ x in 𝓝[<] b, deriv f x ≤ 0) (h₁ : ∀ᶠ x in 𝓝[>] b, deriv f x ≥ 0) : + IsLocalMin f b := by + obtain ⟨a,ha⟩ := (nhdsWithin_Iio_basis' ⟨b - 1, sub_one_lt b⟩).eventually_iff.mp <| hd₀.and h₀ + obtain ⟨c,hc⟩ := (nhdsWithin_Ioi_basis' ⟨b + 1, lt_add_one b⟩).eventually_iff.mp <| hd₁.and h₁ + exact isLocalMin_of_deriv_Ioo ha.1 hc.1 h + (fun _ hx => (ha.2 hx).1.differentiableWithinAt) + (fun _ hx => (hc.2 hx).1.differentiableWithinAt) + (fun _ hx => (ha.2 hx).2) (fun x hx => (hc.2 hx).2) + +/-- The First Derivative test, maximum version. -/ +theorem isLocalMax_of_deriv {f : ℝ → ℝ} {b : ℝ} (h : ContinuousAt f b) + (hd : ∀ᶠ x in 𝓝[≠] b, DifferentiableAt ℝ f x) + (h₀ : ∀ᶠ x in 𝓝[<] b, 0 ≤ deriv f x) (h₁ : ∀ᶠ x in 𝓝[>] b, deriv f x ≤ 0) : + IsLocalMax f b := + isLocalMax_of_deriv' h + (nhds_left'_le_nhds_ne _ (by tauto)) (nhds_right'_le_nhds_ne _ (by tauto)) h₀ h₁ + +/-- The First Derivative test, minimum version. -/ +theorem isLocalMin_of_deriv {f : ℝ → ℝ} {b : ℝ} (h : ContinuousAt f b) + (hd : ∀ᶠ x in 𝓝[≠] b, DifferentiableAt ℝ f x) + (h₀ : ∀ᶠ x in 𝓝[<] b, deriv f x ≤ 0) (h₁ : ∀ᶠ x in 𝓝[>] b, 0 ≤ deriv f x) : + IsLocalMin f b := + isLocalMin_of_deriv' h + (nhds_left'_le_nhds_ne _ (by tauto)) (nhds_right'_le_nhds_ne _ (by tauto)) h₀ h₁ diff --git a/Mathlib/Topology/ContinuousOn.lean b/Mathlib/Topology/ContinuousOn.lean index 9cd83acf60189..1de7949f03494 100644 --- a/Mathlib/Topology/ContinuousOn.lean +++ b/Mathlib/Topology/ContinuousOn.lean @@ -205,6 +205,33 @@ theorem nhdsWithin_union (a : α) (s t : Set α) : 𝓝[s ∪ t] a = 𝓝[s] a delta nhdsWithin rw [← inf_sup_left, sup_principal] +theorem nhds_eq_nhdsWithin_sup_nhdsWithin (b : α) {I₁ I₂ : Set α} (hI : Set.univ = I₁ ∪ I₂) : + nhds b = nhdsWithin b I₁ ⊔ nhdsWithin b I₂ := by + rw [← nhdsWithin_univ b, hI, nhdsWithin_union] + +/-- If `L` and `R` are neighborhoods of `b` within sets whose union is `Set.univ`, then +`L ∪ R` is a neighborhood of `b`. -/ +theorem union_mem_nhds_of_mem_nhdsWithin {b : α} + {I₁ I₂ : Set α} (h : Set.univ = I₁ ∪ I₂) + {L : Set α} (hL : L ∈ nhdsWithin b I₁) + {R : Set α} (hR : R ∈ nhdsWithin b I₂) : L ∪ R ∈ nhds b := by + rw [← nhdsWithin_univ b, h, nhdsWithin_union] + exact ⟨mem_of_superset hL (by aesop), mem_of_superset hR (by aesop)⟩ + + +/-- Writing a punctured neighborhood filter as a sup of left and right filters. -/ +lemma punctured_nhds_eq_nhdsWithin_sup_nhdsWithin [LinearOrder α] {x : α} : + 𝓝[≠] x = 𝓝[<] x ⊔ 𝓝[>] x := by + rw [← Iio_union_Ioi, nhdsWithin_union] + + +/-- Obtain a "predictably-sided" neighborhood of `b` from two one-sided neighborhoods. -/ +theorem nhds_of_Ici_Iic [LinearOrder α] {b : α} + {L : Set α} (hL : L ∈ 𝓝[≤] b) + {R : Set α} (hR : R ∈ 𝓝[≥] b) : L ∩ Iic b ∪ R ∩ Ici b ∈ 𝓝 b := + union_mem_nhds_of_mem_nhdsWithin Iic_union_Ici.symm + (inter_mem hL self_mem_nhdsWithin) (inter_mem hR self_mem_nhdsWithin) + theorem nhdsWithin_biUnion {ι} {I : Set ι} (hI : I.Finite) (s : ι → Set α) (a : α) : 𝓝[⋃ i ∈ I, s i] a = ⨆ i ∈ I, 𝓝[s i] a := Set.Finite.induction_on hI (by simp) fun _ _ hT ↦ by @@ -1278,3 +1305,15 @@ theorem continuousOn_piecewise_ite {s s' t : Set α} {f f' : α → β} [∀ x, (h : ContinuousOn f s) (h' : ContinuousOn f' s') (H : s ∩ frontier t = s' ∩ frontier t) (Heq : EqOn f f' (s ∩ frontier t)) : ContinuousOn (t.piecewise f f') (t.ite s s') := continuousOn_piecewise_ite' (h.mono inter_subset_left) (h'.mono inter_subset_left) H Heq + + +/-- If `f` is continuous on an open set `s` and continuous at each point of another +set `t` then `f` is continuous on `s ∪ t`. -/ +lemma ContinuousOn.union_continuousAt + {X Y : Type*} [TopologicalSpace X] [TopologicalSpace Y] + {s t : Set X} {f : X → Y} (s_op : IsOpen s) + (hs : ContinuousOn f s) (ht : ∀ x ∈ t, ContinuousAt f x) : + ContinuousOn f (s ∪ t) := + ContinuousAt.continuousOn <| fun _ hx => hx.elim + (fun h => ContinuousWithinAt.continuousAt (continuousWithinAt hs h) <| IsOpen.mem_nhds s_op h) + (ht _) diff --git a/Mathlib/Topology/Order/LocalExtr.lean b/Mathlib/Topology/Order/LocalExtr.lean index 8253ac222956f..6e58a0c6754b9 100644 --- a/Mathlib/Topology/Order/LocalExtr.lean +++ b/Mathlib/Topology/Order/LocalExtr.lean @@ -520,3 +520,21 @@ theorem Filter.EventuallyEq.isLocalExtr_iff {f g : α → β} {a : α} (heq : f heq.isExtrFilter_iff heq.eq_of_nhds end Eventually + +/-- If `f` is monotone to the left and antitone to the right, then it has a local maximum. -/ +lemma isLocalMax_of_mono_anti' {α : Type*} [TopologicalSpace α] [LinearOrder α] + {β : Type*} [Preorder β] {b : α} {f : α → β} + {a : Set α} (ha : a ∈ 𝓝[≤] b) {c : Set α} (hc : c ∈ 𝓝[≥] b) + (h₀ : MonotoneOn f a) (h₁ : AntitoneOn f c) : IsLocalMax f b := + have : b ∈ a := mem_of_mem_nhdsWithin (by simp) ha + have : b ∈ c := mem_of_mem_nhdsWithin (by simp) hc + mem_of_superset (nhds_of_Ici_Iic ha hc) (fun x _ => by rcases le_total x b <;> aesop) + +/-- If `f` is antitone to the left and monotone to the right, then it has a local minimum. -/ +lemma isLocalMin_of_anti_mono' {α : Type*} [TopologicalSpace α] [LinearOrder α] + {β : Type*} [Preorder β] {b : α} {f : α → β} + {a : Set α} (ha : a ∈ 𝓝[≤] b) {c : Set α} (hc : c ∈ 𝓝[≥] b) + (h₀ : AntitoneOn f a) (h₁ : MonotoneOn f c) : IsLocalMin f b := + have : b ∈ a := mem_of_mem_nhdsWithin (by simp) ha + have : b ∈ c := mem_of_mem_nhdsWithin (by simp) hc + mem_of_superset (nhds_of_Ici_Iic ha hc) (fun x _ => by rcases le_total x b <;> aesop) diff --git a/Mathlib/Topology/Order/OrderClosedExtr.lean b/Mathlib/Topology/Order/OrderClosedExtr.lean new file mode 100644 index 0000000000000..0e6754114a397 --- /dev/null +++ b/Mathlib/Topology/Order/OrderClosedExtr.lean @@ -0,0 +1,49 @@ +/- +Copyright (c) 2024 Bjørn Kjos-Hanssen. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Bjørn Kjos-Hanssen, Patrick Massot +-/ + +import Mathlib.Topology.Order.OrderClosed +import Mathlib.Topology.Order.LocalExtr + +/-! +# Local maxima from monotonicity and antitonicity + +In this file we prove a lemma that is useful for the First Derivative Test in calculus, +and its dual. + +## Main statements + +* `isLocalMax_of_mono_anti` : if a function `f` is monotone to the left of `x` + and antitone to the right of `x` then `f` has a local maximum at `x`. + +* `isLocalMin_of_anti_mono` : the dual statement for minima. + +* `isLocalMax_of_mono_anti'` : a version of `isLocalMax_of_mono_anti` for filters. + +* `isLocalMin_of_anti_mono'` : a version of `isLocalMax_of_mono_anti'` for minima. + +-/ + +open Set Topology Filter + +/-- If `f` is monotone on `(a,b]` and antitone on `[b,c)` then `f` has +a local maximum at `b`. -/ +lemma isLocalMax_of_mono_anti + {α : Type*} [TopologicalSpace α] [LinearOrder α] [OrderClosedTopology α] + {β : Type*} [Preorder β] + {a b c : α} (g₀ : a < b) (g₁ : b < c) {f : α → β} + (h₀ : MonotoneOn f (Ioc a b)) + (h₁ : AntitoneOn f (Ico b c)) : IsLocalMax f b := + isLocalMax_of_mono_anti' (Ioc_mem_nhdsWithin_Iic' g₀) (Ico_mem_nhdsWithin_Ici' g₁) h₀ h₁ + + + +/-- If `f` is antitone on `(a,b]` and monotone on `[b,c)` then `f` has +a local minimum at `b`. -/ +lemma isLocalMin_of_anti_mono + {α : Type*} [TopologicalSpace α] [LinearOrder α] [OrderClosedTopology α] + {β : Type*} [Preorder β] {a b c : α} (g₀ : a < b) (g₁ : b < c) {f : α → β} + (h₀ : AntitoneOn f (Ioc a b)) (h₁ : MonotoneOn f (Ico b c)) : IsLocalMin f b := + mem_of_superset (Ioo_mem_nhds g₀ g₁) (fun x hx => by rcases le_total x b <;> aesop) From ed8048704cc22467cc23c9796e4b83e4d79edda8 Mon Sep 17 00:00:00 2001 From: Bhavik Mehta Date: Thu, 24 Oct 2024 00:12:10 +0000 Subject: [PATCH 353/505] feat(Matrix): extracting entries as a linear map (#17535) Strictly speaking, something like this was available as `LinearMap.proj`, but actually using it in practice is awkward as unification struggles (see line 1227 in this PR). I'm happy to take name suggestions, as well as ideas for other generalities. Co-authored-by: Eric Wieser --- Mathlib/Data/Matrix/Basic.lean | 127 +++++++++++++++++++++++++++++++++ 1 file changed, 127 insertions(+) diff --git a/Mathlib/Data/Matrix/Basic.lean b/Mathlib/Data/Matrix/Basic.lean index 6194d386d0f50..b329c00a298d1 100644 --- a/Mathlib/Data/Matrix/Basic.lean +++ b/Mathlib/Data/Matrix/Basic.lean @@ -14,6 +14,7 @@ import Mathlib.Algebra.Star.BigOperators import Mathlib.Algebra.Star.Module import Mathlib.Algebra.Star.Pi import Mathlib.Data.Fintype.BigOperators +import Mathlib.LinearAlgebra.Pi /-! # Matrices @@ -334,6 +335,26 @@ instance subsingleton_of_empty_right [IsEmpty n] : Subsingleton (Matrix m n α) ext i j exact isEmptyElim j⟩ +/-- This is `Matrix.of` bundled as an additive equivalence. -/ +def ofAddEquiv [Add α] : (m → n → α) ≃+ Matrix m n α where + __ := of + map_add' _ _ := rfl + +@[simp] lemma coe_ofAddEquiv [Add α] : + ⇑(ofAddEquiv : (m → n → α) ≃+ Matrix m n α) = of := rfl +@[simp] lemma coe_ofAddEquiv_symm [Add α] : + ⇑(ofAddEquiv.symm : Matrix m n α ≃+ (m → n → α)) = of.symm := rfl + +/-- This is `Matrix.of` bundled as a linear equivalence. -/ +def ofLinearEquiv [Semiring R] [AddCommMonoid α] [Module R α] : (m → n → α) ≃ₗ[R] Matrix m n α where + __ := ofAddEquiv + map_smul' _ _ := rfl + +@[simp] lemma coe_ofLinearEquiv [Semiring R] [AddCommMonoid α] [Module R α] : + ⇑(ofLinearEquiv : (m → n → α) ≃ₗ[R] Matrix m n α) = of := rfl +@[simp] lemma coe_ofLinearEquiv_symm [Semiring R] [AddCommMonoid α] [Module R α] : + ⇑(ofLinearEquiv.symm : Matrix m n α ≃ₗ[R] (m → n → α)) = of.symm := rfl + end Matrix open Matrix @@ -1209,6 +1230,93 @@ def diagonalAlgHom : (n → α) →ₐ[R] Matrix n n α := end Algebra +section AddHom + +variable [Add α] + +variable (R α) in +/-- Extracting entries from a matrix as an additive homomorphism. -/ +@[simps] +def entryAddHom (i : m) (j : n) : AddHom (Matrix m n α) α where + toFun M := M i j + map_add' _ _ := rfl + +-- It is necessary to spell out the name of the coercion explicitly on the RHS +-- for unification to succeed +lemma entryAddHom_eq_comp {i : m} {j : n} : + entryAddHom α i j = + ((Pi.evalAddHom _ j).comp (Pi.evalAddHom _ i)).comp (AddHomClass.toAddHom ofAddEquiv.symm) := + rfl + +end AddHom + +section AddMonoidHom + +variable [AddZeroClass α] + +variable (R α) in +/-- +Extracting entries from a matrix as an additive monoid homomorphism. Note this cannot be upgraded to +a ring homomorphism, as it does not respect multiplication. +-/ +@[simps] +def entryAddMonoidHom (i : m) (j : n) : Matrix m n α →+ α where + toFun M := M i j + map_add' _ _ := rfl + map_zero' := rfl + +-- It is necessary to spell out the name of the coercion explicitly on the RHS +-- for unification to succeed +lemma entryAddMonoidHom_eq_comp {i : m} {j : n} : + entryAddMonoidHom α i j = + ((Pi.evalAddMonoidHom _ j).comp (Pi.evalAddMonoidHom _ i)).comp + (AddMonoidHomClass.toAddMonoidHom ofAddEquiv.symm) := by + rfl + +@[simp] lemma evalAddMonoidHom_comp_diagAddMonoidHom (i : m) : + (Pi.evalAddMonoidHom _ i).comp (diagAddMonoidHom m α) = entryAddMonoidHom α i i := by + simp [AddMonoidHom.ext_iff] + +@[simp] lemma entryAddMonoidHom_toAddHom {i : m} {j : n} : + (entryAddMonoidHom α i j : AddHom _ _) = entryAddHom α i j := rfl + +end AddMonoidHom + +section LinearMap + +variable [Semiring R] [AddCommMonoid α] [Module R α] + +variable (R α) in +/-- +Extracting entries from a matrix as a linear map. Note this cannot be upgraded to an algebra +homomorphism, as it does not respect multiplication. +-/ +@[simps] +def entryLinearMap (i : m) (j : n) : + Matrix m n α →ₗ[R] α where + toFun M := M i j + map_add' _ _ := rfl + map_smul' _ _ := rfl + +-- It is necessary to spell out the name of the coercion explicitly on the RHS +-- for unification to succeed +lemma entryLinearMap_eq_comp {i : m} {j : n} : + entryLinearMap R α i j = + LinearMap.proj j ∘ₗ LinearMap.proj i ∘ₗ ofLinearEquiv.symm.toLinearMap := by + rfl + +@[simp] lemma proj_comp_diagLinearMap (i : m) : + LinearMap.proj i ∘ₗ diagLinearMap m R α = entryLinearMap R α i i := by + simp [LinearMap.ext_iff] + +@[simp] lemma entryLinearMap_toAddMonoidHom {i : m} {j : n} : + (entryLinearMap R α i j : _ →+ _) = entryAddMonoidHom α i j := rfl + +@[simp] lemma entryLinearMap_toAddHom {i : m} {j : n} : + (entryLinearMap R α i j : AddHom _ _) = entryAddHom α i j := rfl + +end LinearMap + end Matrix /-! @@ -1263,6 +1371,9 @@ theorem mapMatrix_comp (f : β →+ γ) (g : α →+ β) : f.mapMatrix.comp g.mapMatrix = ((f.comp g).mapMatrix : Matrix m n α →+ _) := rfl +@[simp] lemma entryAddMonoidHom_comp_mapMatrix (f : α →+ β) (i : m) (j : n) : + (entryAddMonoidHom β i j).comp f.mapMatrix = f.comp (entryAddMonoidHom α i j) := rfl + end AddMonoidHom namespace AddEquiv @@ -1291,6 +1402,10 @@ theorem mapMatrix_trans (f : α ≃+ β) (g : β ≃+ γ) : f.mapMatrix.trans g.mapMatrix = ((f.trans g).mapMatrix : Matrix m n α ≃+ _) := rfl +@[simp] lemma entryAddHom_comp_mapMatrix (f : α ≃+ β) (i : m) (j : n) : + (entryAddHom β i j).comp (AddHomClass.toAddHom f.mapMatrix) = + (f : AddHom α β).comp (entryAddHom _ i j) := rfl + end AddEquiv namespace LinearMap @@ -1315,6 +1430,9 @@ theorem mapMatrix_comp (f : β →ₗ[R] γ) (g : α →ₗ[R] β) : f.mapMatrix.comp g.mapMatrix = ((f.comp g).mapMatrix : Matrix m n α →ₗ[R] _) := rfl +@[simp] lemma entryLinearMap_comp_mapMatrix (f : α →ₗ[R] β) (i : m) (j : n) : + entryLinearMap R _ i j ∘ₗ f.mapMatrix = f ∘ₗ entryLinearMap R _ i j := rfl + end LinearMap namespace LinearEquiv @@ -1345,6 +1463,15 @@ theorem mapMatrix_trans (f : α ≃ₗ[R] β) (g : β ≃ₗ[R] γ) : f.mapMatrix.trans g.mapMatrix = ((f.trans g).mapMatrix : Matrix m n α ≃ₗ[R] _) := rfl +@[simp] lemma mapMatrix_toLinearMap (f : α ≃ₗ[R] β) : + (f.mapMatrix : _ ≃ₗ[R] Matrix m n β).toLinearMap = f.toLinearMap.mapMatrix := by + rfl + +@[simp] lemma entryLinearMap_comp_mapMatrix (f : α ≃ₗ[R] β) (i : m) (j : n) : + entryLinearMap R _ i j ∘ₗ f.mapMatrix.toLinearMap = + f.toLinearMap ∘ₗ entryLinearMap R _ i j := by + simp only [mapMatrix_toLinearMap, LinearMap.entryLinearMap_comp_mapMatrix] + end LinearEquiv namespace RingHom From 15e91eba82bde41430e8c72ccf0e96a29337f25b Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 24 Oct 2024 00:56:07 +0000 Subject: [PATCH 354/505] chore: golf a messy Matrix proof (#18097) Using the new `Matrix.ofLinearEquiv` casting function here (modified in this PR to take an explicit argument) makes everything go through cleanly, resolving the porting note. --- Mathlib/Data/Matrix/Basic.lean | 11 ++++++++--- Mathlib/Data/Matrix/Basis.lean | 6 ++++++ Mathlib/LinearAlgebra/StdBasis.lean | 21 ++------------------- 3 files changed, 16 insertions(+), 22 deletions(-) diff --git a/Mathlib/Data/Matrix/Basic.lean b/Mathlib/Data/Matrix/Basic.lean index b329c00a298d1..543197f1454dc 100644 --- a/Mathlib/Data/Matrix/Basic.lean +++ b/Mathlib/Data/Matrix/Basic.lean @@ -345,15 +345,20 @@ def ofAddEquiv [Add α] : (m → n → α) ≃+ Matrix m n α where @[simp] lemma coe_ofAddEquiv_symm [Add α] : ⇑(ofAddEquiv.symm : Matrix m n α ≃+ (m → n → α)) = of.symm := rfl +section +variable (R) + /-- This is `Matrix.of` bundled as a linear equivalence. -/ def ofLinearEquiv [Semiring R] [AddCommMonoid α] [Module R α] : (m → n → α) ≃ₗ[R] Matrix m n α where __ := ofAddEquiv map_smul' _ _ := rfl @[simp] lemma coe_ofLinearEquiv [Semiring R] [AddCommMonoid α] [Module R α] : - ⇑(ofLinearEquiv : (m → n → α) ≃ₗ[R] Matrix m n α) = of := rfl + ⇑(ofLinearEquiv _ : (m → n → α) ≃ₗ[R] Matrix m n α) = of := rfl @[simp] lemma coe_ofLinearEquiv_symm [Semiring R] [AddCommMonoid α] [Module R α] : - ⇑(ofLinearEquiv.symm : Matrix m n α ≃ₗ[R] (m → n → α)) = of.symm := rfl + ⇑((ofLinearEquiv _).symm : Matrix m n α ≃ₗ[R] (m → n → α)) = of.symm := rfl + +end end Matrix @@ -1302,7 +1307,7 @@ def entryLinearMap (i : m) (j : n) : -- for unification to succeed lemma entryLinearMap_eq_comp {i : m} {j : n} : entryLinearMap R α i j = - LinearMap.proj j ∘ₗ LinearMap.proj i ∘ₗ ofLinearEquiv.symm.toLinearMap := by + LinearMap.proj j ∘ₗ LinearMap.proj i ∘ₗ (ofLinearEquiv R).symm.toLinearMap := by rfl @[simp] lemma proj_comp_diagLinearMap (i : m) : diff --git a/Mathlib/Data/Matrix/Basis.lean b/Mathlib/Data/Matrix/Basis.lean index 01c7a8cf445b6..4b4af6da0eab3 100644 --- a/Mathlib/Data/Matrix/Basis.lean +++ b/Mathlib/Data/Matrix/Basis.lean @@ -30,6 +30,12 @@ and zeroes elsewhere. def stdBasisMatrix (i : m) (j : n) (a : α) : Matrix m n α := fun i' j' => if i = i' ∧ j = j' then a else 0 +theorem stdBasisMatrix_eq_of_single_single (i : m) (j : n) (a : α) : + stdBasisMatrix i j a = Matrix.of (Pi.single i (Pi.single j a)) := by + ext a b + unfold stdBasisMatrix + by_cases hi : i = a <;> by_cases hj : j = b <;> simp [*] + @[simp] theorem smul_stdBasisMatrix [SMulZeroClass R α] (r : R) (i : m) (j : n) (a : α) : r • stdBasisMatrix i j a = stdBasisMatrix i j (r • a) := by diff --git a/Mathlib/LinearAlgebra/StdBasis.lean b/Mathlib/LinearAlgebra/StdBasis.lean index bd602506eb518..3159915328c21 100644 --- a/Mathlib/LinearAlgebra/StdBasis.lean +++ b/Mathlib/LinearAlgebra/StdBasis.lean @@ -294,29 +294,12 @@ variable (R : Type*) (m n : Type*) [Fintype m] [Finite n] [Semiring R] /-- The standard basis of `Matrix m n R`. -/ noncomputable def stdBasis : Basis (m × n) R (Matrix m n R) := Basis.reindex (Pi.basis fun _ : m => Pi.basisFun R n) (Equiv.sigmaEquivProd _ _) + |>.map (ofLinearEquiv R) variable {n m} theorem stdBasis_eq_stdBasisMatrix (i : m) (j : n) [DecidableEq m] [DecidableEq n] : stdBasis R m n (i, j) = stdBasisMatrix i j (1 : R) := by - -- Porting note: `simp` fails to apply `Pi.basis_apply` - ext a b - by_cases hi : i = a <;> by_cases hj : j = b - · simp only [stdBasis, hi, hj, Basis.coe_reindex, comp_apply, Equiv.sigmaEquivProd_symm_apply, - StdBasisMatrix.apply_same] - erw [Pi.basis_apply] - simp - · simp only [stdBasis, hi, Basis.coe_reindex, comp_apply, Equiv.sigmaEquivProd_symm_apply, - hj, and_false, not_false_iff, StdBasisMatrix.apply_of_ne] - erw [Pi.basis_apply] - simp [hj] - · simp only [stdBasis, hj, Basis.coe_reindex, comp_apply, Equiv.sigmaEquivProd_symm_apply, - hi, and_true, not_false_iff, StdBasisMatrix.apply_of_ne] - erw [Pi.basis_apply] - simp [hi, hj, Ne.symm hi, Pi.single_eq_of_ne] - · simp only [stdBasis, Basis.coe_reindex, comp_apply, Equiv.sigmaEquivProd_symm_apply, - hi, hj, and_self, not_false_iff, StdBasisMatrix.apply_of_ne] - erw [Pi.basis_apply] - simp [hi, hj, Ne.symm hj, Ne.symm hi, Pi.single_eq_of_ne] + simp [stdBasis, stdBasisMatrix_eq_of_single_single] end Matrix From 1f4900e1df2e7dc98c86059444e2f03608dd4bfc Mon Sep 17 00:00:00 2001 From: Markus Himmel Date: Thu, 24 Oct 2024 04:37:46 +0000 Subject: [PATCH 355/505] feat(CategoryTheory): translate between Yoneda and "relative Yoneda" (#17449) --- Mathlib/CategoryTheory/Comma/Presheaf.lean | 69 ++++++++++++++++++++++ 1 file changed, 69 insertions(+) diff --git a/Mathlib/CategoryTheory/Comma/Presheaf.lean b/Mathlib/CategoryTheory/Comma/Presheaf.lean index 8fa1a25d073e5..56df6eace290c 100644 --- a/Mathlib/CategoryTheory/Comma/Presheaf.lean +++ b/Mathlib/CategoryTheory/Comma/Presheaf.lean @@ -3,6 +3,7 @@ Copyright (c) 2024 Markus Himmel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Markus Himmel -/ +import Mathlib.CategoryTheory.HomCongr import Mathlib.CategoryTheory.Comma.Over import Mathlib.Tactic.CategoryTheory.Elementwise @@ -572,4 +573,72 @@ def CostructuredArrow.toOverCompOverEquivPresheafCostructuredArrow (A : Cᵒᵖ CostructuredArrow.toOver yoneda A ⋙ (overEquivPresheafCostructuredArrow A).functor ≅ yoneda := toOverYonedaCompRestrictedYoneda A +/-- This isomorphism says that hom-sets in the category `Over A` for a presheaf `A` where the domain + is of the form `(CostructuredArrow.toOver yoneda A).obj X` can instead be interpreted as + hom-sets in the category `(CostructuredArrow yoneda A)ᵒᵖ ⥤ Type v` where the domain is of the + form `yoneda.obj X` after adjusting the codomain accordingly. This is desirable because in the + latter case the Yoneda lemma can be applied. -/ +def CostructuredArrow.toOverCompYoneda (A : Cᵒᵖ ⥤ Type v) (T : Over A) : + (CostructuredArrow.toOver yoneda A).op ⋙ yoneda.obj T ≅ + yoneda.op ⋙ yoneda.obj ((overEquivPresheafCostructuredArrow A).functor.obj T) := + NatIso.ofComponents (fun X => + (overEquivPresheafCostructuredArrow A).fullyFaithfulFunctor.homEquiv.toIso ≪≫ + (Iso.homCongr + ((CostructuredArrow.toOverCompOverEquivPresheafCostructuredArrow A).app X.unop) + (Iso.refl _)).toIso) + (by aesop_cat) + +@[simp] +theorem CostructuredArrow.overEquivPresheafCostructuredArrow_inverse_map_toOverCompYoneda + {A : Cᵒᵖ ⥤ Type v} {T : Over A} {X : CostructuredArrow yoneda A} + (f : (CostructuredArrow.toOver yoneda A).obj X ⟶ T) : + (overEquivPresheafCostructuredArrow A).inverse.map + (((CostructuredArrow.toOverCompYoneda A T).hom.app (op X) f)) = + (CostructuredArrow.toOverCompOverEquivPresheafCostructuredArrow A).isoCompInverse.inv.app X ≫ + f ≫ (overEquivPresheafCostructuredArrow A).unit.app T := by + simp [CostructuredArrow.toOverCompYoneda] + +@[simp] +theorem CostructuredArrow.overEquivPresheafCostructuredArrow_functor_map_toOverCompYoneda + {A : Cᵒᵖ ⥤ Type v} {T : Over A} {X : CostructuredArrow yoneda A} + (f : yoneda.obj X ⟶ (overEquivPresheafCostructuredArrow A).functor.obj T) : + (overEquivPresheafCostructuredArrow A).functor.map + (((CostructuredArrow.toOverCompYoneda A T).inv.app (op X) f)) = + (CostructuredArrow.toOverCompOverEquivPresheafCostructuredArrow A).hom.app X ≫ f := by + simp [CostructuredArrow.toOverCompYoneda] + +/-- This isomorphism says that hom-sets in the category `Over A` for a presheaf `A` where the domain + is of the form `(CostructuredArrow.toOver yoneda A).obj X` can instead be interpreted as + hom-sets in the category `(CostructuredArrow yoneda A)ᵒᵖ ⥤ Type v` where the domain is of the + form `yoneda.obj X` after adjusting the codomain accordingly. This is desirable because in the + latter case the Yoneda lemma can be applied. -/ +def CostructuredArrow.toOverCompCoyoneda (A : Cᵒᵖ ⥤ Type v) : + (CostructuredArrow.toOver yoneda A).op ⋙ coyoneda ≅ + yoneda.op ⋙ coyoneda ⋙ + (whiskeringLeft _ _ _).obj (overEquivPresheafCostructuredArrow A).functor := + NatIso.ofComponents (fun X => NatIso.ofComponents (fun Y => + (overEquivPresheafCostructuredArrow A).fullyFaithfulFunctor.homEquiv.toIso ≪≫ + (Iso.homCongr + ((CostructuredArrow.toOverCompOverEquivPresheafCostructuredArrow A).app X.unop) + (Iso.refl _)).toIso)) (by aesop_cat) + +@[simp] +theorem CostructuredArrow.overEquivPresheafCostructuredArrow_inverse_map_toOverCompCoyoneda + {A : Cᵒᵖ ⥤ Type v} {T : Over A} {X : CostructuredArrow yoneda A} + (f : (CostructuredArrow.toOver yoneda A).obj X ⟶ T) : + (overEquivPresheafCostructuredArrow A).inverse.map + (((CostructuredArrow.toOverCompCoyoneda A).hom.app (op X)).app T f) = + (CostructuredArrow.toOverCompOverEquivPresheafCostructuredArrow A).isoCompInverse.inv.app X ≫ + f ≫ (overEquivPresheafCostructuredArrow A).unit.app T := by + simp [CostructuredArrow.toOverCompCoyoneda] + +@[simp] +theorem CostructuredArrow.overEquivPresheafCostructuredArrow_functor_map_toOverCompCoyoneda + {A : Cᵒᵖ ⥤ Type v} {T : Over A} {X : CostructuredArrow yoneda A} + (f : yoneda.obj X ⟶ (overEquivPresheafCostructuredArrow A).functor.obj T) : + (overEquivPresheafCostructuredArrow A).functor.map + (((CostructuredArrow.toOverCompCoyoneda A).inv.app (op X)).app T f) = + (CostructuredArrow.toOverCompOverEquivPresheafCostructuredArrow A).hom.app X ≫ f := by + simp [CostructuredArrow.toOverCompCoyoneda] + end CategoryTheory From a5464da2f70c07f829592cdc8c1befc1cb65dfb8 Mon Sep 17 00:00:00 2001 From: Christopher Hoskin Date: Thu, 24 Oct 2024 04:49:46 +0000 Subject: [PATCH 356/505] =?UTF-8?q?feat:=20`LinearEquiv.congrQuadraticMap`?= =?UTF-8?q?=20and=20`LinearEquiv.congrRight=E2=82=82`=20(#17454)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit `LinearEquiv.congrQuadraticMap` matches `LinearMap.compQuadraticMap`, while `LinearEquiv.congrRight₂` somewhat matches `LinearMap.compr₂`. Used in #14988 Co-authored-by: Christopher Hoskin Co-authored-by: Eric Wieser --- Mathlib/LinearAlgebra/BilinearForm/Hom.lean | 32 +++++++++++++++++++ .../LinearAlgebra/QuadraticForm/Basic.lean | 22 +++++++++++++ 2 files changed, 54 insertions(+) diff --git a/Mathlib/LinearAlgebra/BilinearForm/Hom.lean b/Mathlib/LinearAlgebra/BilinearForm/Hom.lean index 98d0c01aeaba5..17f78fa9dc6ea 100644 --- a/Mathlib/LinearAlgebra/BilinearForm/Hom.lean +++ b/Mathlib/LinearAlgebra/BilinearForm/Hom.lean @@ -39,6 +39,7 @@ Bilinear form, -/ open LinearMap (BilinForm) +open LinearMap (BilinMap) universe u v w @@ -307,6 +308,37 @@ theorem comp_congr (e : M' ≃ₗ[R] M'') (B : BilinForm R M) (l r : M' →ₗ[R end congr +section congrRight₂ + +variable {N₁ N₂ N₃ : Type*} +variable [AddCommMonoid N₁] [AddCommMonoid N₂] [AddCommMonoid N₃] +variable [Module R N₁] [Module R N₂] [Module R N₃] + +/-- When `N₁` and `N₂` are equivalent, bilinear maps on `M` into `N₁` are equivalent to bilinear +maps into `N₂`. -/ +def _root_.LinearEquiv.congrRight₂ (e : N₁ ≃ₗ[R] N₂) : BilinMap R M N₁ ≃ₗ[R] BilinMap R M N₂ := + LinearEquiv.congrRight (LinearEquiv.congrRight e) + +@[simp] +theorem _root_.LinearEquiv.congrRight₂_apply (e : N₁ ≃ₗ[R] N₂) (B : BilinMap R M N₁) : + LinearEquiv.congrRight₂ e B = compr₂ B e := rfl + +@[simp] +theorem _root_.LinearEquiv.congrRight₂_refl : + LinearEquiv.congrRight₂ (.refl R N₁) = .refl R (BilinMap R M N₁) := rfl + +@[simp] +theorem _root_.LinearEquiv.congrRight_symm (e : N₁ ≃ₗ[R] N₂) : + (LinearEquiv.congrRight₂ e (M := M)).symm = LinearEquiv.congrRight₂ e.symm := + rfl + +theorem _root_.LinearEquiv.congrRight₂_trans (e₁₂ : N₁ ≃ₗ[R] N₂) (e₂₃ : N₂ ≃ₗ[R] N₃) : + LinearEquiv.congrRight₂ (M := M) (e₁₂ ≪≫ₗ e₂₃) = + LinearEquiv.congrRight₂ e₁₂ ≪≫ₗ LinearEquiv.congrRight₂ e₂₃ := + rfl + +end congrRight₂ + section LinMulLin /-- `linMulLin f g` is the bilinear form mapping `x` and `y` to `f x * g y` -/ diff --git a/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean b/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean index 99076960053fe..41254cc0f29bb 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean @@ -574,6 +574,28 @@ def _root_.LinearMap.compQuadraticMap' [CommSemiring S] [Algebra S R] [Module S (f : N →ₗ[S] P) (Q : QuadraticMap R M N) : QuadraticMap S M P := _root_.LinearMap.compQuadraticMap f Q.restrictScalars +/-- When `N` and `P` are equivalent, quadratic maps on `M` into `N` are equivalent to quadratic +maps on `M` into `P`. + +See `LinearMap.BilinMap.congr₂` for the bilinear map version. -/ +@[simps] +def _root_.LinearEquiv.congrQuadraticMap (e : N ≃ₗ[R] P) : + QuadraticMap R M N ≃ₗ[R] QuadraticMap R M P where + toFun Q := e.compQuadraticMap Q + invFun Q := e.symm.compQuadraticMap Q + left_inv _ := ext fun _ => e.symm_apply_apply _ + right_inv _ := ext fun _ => e.apply_symm_apply _ + map_add' _ _ := ext fun _ => map_add e _ _ + map_smul' _ _ := ext fun _ => _root_.map_smul e _ _ + +@[simp] +theorem _root_.LinearEquiv.congrQuadraticMap_refl : + LinearEquiv.congrQuadraticMap (.refl R N) = .refl R (QuadraticMap R M N) := rfl + +@[simp] +theorem _root_.LinearEquiv.congrQuadraticMap_symm (e : N ≃ₗ[R] P) : + (LinearEquiv.congrQuadraticMap e (M := M)).symm = e.symm.congrQuadraticMap := rfl + end Comp section NonUnitalNonAssocSemiring From c46a2f187c308d6e4523de78efd44f1b9824b549 Mon Sep 17 00:00:00 2001 From: Christian Merten Date: Thu, 24 Oct 2024 04:49:47 +0000 Subject: [PATCH 357/505] feat(CategoryTheory/Comma): subcategory defined by morphism properties (#18088) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Given functors `L : A ⥤ T` and `R : B ⥤ T` and morphism properties `P`, `Q` and `W` on `T`, `A` and `B` respectively, we define the subcategory `P.Comma L R Q W` of `Comma L R` where - objects are objects of `Comma L R` with the structural morphism satisfying `P`, and - morphisms are morphisms of `Comma L R` where the left morphism satisfies `Q` and the right morphism satisfies `W`. For an object `X : T`, this specializes to `P.Over Q X` which is the subcategory of `Over X` where the structural morphism satisfies `P` and where the horizontal morphisms satisfy `Q`. Common examples of the latter are e.g. the category of schemes étale (finite, affine, etc.) over a base `X`. Here `Q = ⊤`. Co-authored-by: Christian Merten <136261474+chrisflav@users.noreply.github.com> --- Mathlib.lean | 1 + Mathlib/CategoryTheory/Comma/Basic.lean | 28 ++ .../MorphismProperty/Comma.lean | 243 ++++++++++++++++++ .../MorphismProperty/Composition.lean | 4 + 4 files changed, 276 insertions(+) create mode 100644 Mathlib/CategoryTheory/MorphismProperty/Comma.lean diff --git a/Mathlib.lean b/Mathlib.lean index e9a3b499c9093..488148a28f8e6 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1830,6 +1830,7 @@ import Mathlib.CategoryTheory.Monoidal.Types.Basic import Mathlib.CategoryTheory.Monoidal.Types.Coyoneda import Mathlib.CategoryTheory.Monoidal.Types.Symmetric import Mathlib.CategoryTheory.MorphismProperty.Basic +import Mathlib.CategoryTheory.MorphismProperty.Comma import Mathlib.CategoryTheory.MorphismProperty.Composition import Mathlib.CategoryTheory.MorphismProperty.Concrete import Mathlib.CategoryTheory.MorphismProperty.Factorization diff --git a/Mathlib/CategoryTheory/Comma/Basic.lean b/Mathlib/CategoryTheory/Comma/Basic.lean index 192ef0e4606e0..01cb2d222b041 100644 --- a/Mathlib/CategoryTheory/Comma/Basic.lean +++ b/Mathlib/CategoryTheory/Comma/Basic.lean @@ -163,6 +163,34 @@ theorem eqToHom_right (X Y : Comma L R) (H : X = Y) : section +variable {X Y : Comma L R} (e : X ⟶ Y) + +instance [IsIso e] : IsIso e.left := + (Comma.fst L R).map_isIso e + +instance [IsIso e] : IsIso e.right := + (Comma.snd L R).map_isIso e + +@[simp] +lemma inv_left [IsIso e] : (inv e).left = inv e.left := by + apply IsIso.eq_inv_of_hom_inv_id + rw [← Comma.comp_left, IsIso.hom_inv_id, id_left] + +@[simp] +lemma inv_right [IsIso e] : (inv e).right = inv e.right := by + apply IsIso.eq_inv_of_hom_inv_id + rw [← Comma.comp_right, IsIso.hom_inv_id, id_right] + +lemma left_hom_inv_right [IsIso e] : L.map (e.left) ≫ Y.hom ≫ R.map (inv e.right) = X.hom := by + simp + +lemma inv_left_hom_right [IsIso e] : L.map (inv e.left) ≫ X.hom ≫ R.map e.right = Y.hom := by + simp + +end + +section + variable {L₁ L₂ L₃ : A ⥤ T} {R₁ R₂ R₃ : B ⥤ T} /-- Extract the isomorphism between the left objects from an isomorphism in the comma category. -/ diff --git a/Mathlib/CategoryTheory/MorphismProperty/Comma.lean b/Mathlib/CategoryTheory/MorphismProperty/Comma.lean new file mode 100644 index 0000000000000..605734da78131 --- /dev/null +++ b/Mathlib/CategoryTheory/MorphismProperty/Comma.lean @@ -0,0 +1,243 @@ +/- +Copyright (c) 2024 Christian Merten. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Christian Merten +-/ +import Mathlib.CategoryTheory.Comma.Over +import Mathlib.CategoryTheory.MorphismProperty.Composition + +/-! +# Subcategories of comma categories defined by morphism properties + +Given functors `L : A ⥤ T` and `R : B ⥤ T` and morphism properties `P`, `Q` and `W` +on `T`, A` and `B` respectively, we define the subcategory `P.Comma L R Q W` of +`Comma L R` where + +- objects are objects of `Comma L R` with the structural morphism satisfying `P`, and +- morphisms are morphisms of `Comma L R` where the left morphism satisfies `Q` and the + right morphism satisfies `W`. + +For an object `X : T`, this specializes to `P.Over Q X` which is the subcategory of `Over X` +where the structural morphism satisfies `P` and where the horizontal morphisms satisfy `Q`. +Common examples of the latter are e.g. the category of schemes étale (finite, affine, etc.) +over a base `X`. Here `Q = ⊤`. + +## Implementation details + +- We provide the general constructor `P.Comma L R Q W` to obtain `Over X` and `Under X` as + special cases of the more general setup. + +- Most results are developed only in the case where `Q = ⊤` and `W = ⊤`, but the definition + is setup in the general case to allow for a later generalization if needed. + +-/ + +namespace CategoryTheory.MorphismProperty + +open Limits + +section Comma + +variable {A : Type*} [Category A] {B : Type*} [Category B] {T : Type*} [Category T] + (L : A ⥤ T) (R : B ⥤ T) + +variable (P : MorphismProperty T) (Q : MorphismProperty A) (W : MorphismProperty B) + +/-- `P.Comma L R Q W` is the subcategory of `Comma L R` consisting of +objects `X : Comma L R` where `X.hom` satisfies `P`. The morphisms are given by +morphisms in `Comma L R` where the left one satisfies `Q` and the right one satisfies `W`. -/ +@[ext] +protected structure Comma (Q : MorphismProperty A) (W : MorphismProperty B) extends Comma L R where + prop : P toComma.hom + +namespace Comma + +variable {L R P Q W} + +/-- A morphism in `P.Comma L R Q W` is a morphism in `Comma L R` where the left +hom satisfies `Q` and the right one satisfies `W`. -/ +@[ext] +structure Hom (X Y : P.Comma L R Q W) extends CommaMorphism X.toComma Y.toComma where + prop_hom_left : Q toCommaMorphism.left + prop_hom_right : W toCommaMorphism.right + +/-- The underlying morphism of objects in `Comma L R`. -/ +abbrev Hom.hom {X Y : P.Comma L R Q W} (f : Comma.Hom X Y) : X.toComma ⟶ Y.toComma := + f.toCommaMorphism + +@[simp, nolint simpVarHead] +lemma Hom.hom_mk {X Y : P.Comma L R Q W} + (f : CommaMorphism X.toComma Y.toComma) (hf) (hg) : + Comma.Hom.hom ⟨f, hf, hg⟩ = f := rfl + +/-- See Note [custom simps projection] -/ +def Hom.Simps.hom {X Y : P.Comma L R Q W} (f : X.Hom Y) : + X.toComma ⟶ Y.toComma := + f.hom + +initialize_simps_projections Comma.Hom (toCommaMorphism → hom) + +/-- The identity morphism of an object in `P.Comma L R Q W`. -/ +@[simps] +def id [Q.ContainsIdentities] [W.ContainsIdentities] (X : P.Comma L R Q W) : Comma.Hom X X where + left := 𝟙 X.left + prop_hom_left := Q.id_mem X.toComma.left + prop_hom_right := W.id_mem X.toComma.right + +/-- Composition of morphisms in `P.Comma L R Q W`. -/ +@[simps] +def Hom.comp [Q.IsStableUnderComposition] [W.IsStableUnderComposition] {X Y Z : P.Comma L R Q W} + (f : Comma.Hom X Y) (g : Comma.Hom Y Z) : + Comma.Hom X Z where + left := f.left ≫ g.left + right := f.right ≫ g.right + prop_hom_left := Q.comp_mem _ _ f.prop_hom_left g.prop_hom_left + prop_hom_right := W.comp_mem _ _ f.prop_hom_right g.prop_hom_right + +variable [Q.IsMultiplicative] [W.IsMultiplicative] + +variable (L R P Q W) in +instance : Category (P.Comma L R Q W) where + Hom X Y := X.Hom Y + id X := X.id + comp f g := f.comp g + +/-- Alternative `ext` lemma for `Comma.Hom`. -/ +@[ext] +lemma Hom.ext' {X Y : P.Comma L R Q W} {f g : X ⟶ Y} (h : f.hom = g.hom) : + f = g := Comma.Hom.ext + (congrArg CommaMorphism.left h) + (congrArg CommaMorphism.right h) + +@[simp] +lemma id_hom (X : P.Comma L R Q W) : (𝟙 X : X ⟶ X).hom = 𝟙 X.toComma := rfl + +@[simp] +lemma comp_hom {X Y Z : P.Comma L R Q W} (f : X ⟶ Y) (g : Y ⟶ Z) : + (f ≫ g).hom = f.hom ≫ g.hom := rfl + +/-- If `i` is an isomorphism in `Comma L R`, it is also a morphism in `P.Comma L R Q W`. -/ +@[simps hom] +def homFromCommaOfIsIso [Q.RespectsIso] [W.RespectsIso] {X Y : P.Comma L R Q W} + (i : X.toComma ⟶ Y.toComma) [IsIso i] : + X ⟶ Y where + __ := i + prop_hom_left := Q.of_isIso i.left + prop_hom_right := W.of_isIso i.right + +instance [Q.RespectsIso] [W.RespectsIso] {X Y : P.Comma L R Q W} (i : X.toComma ⟶ Y.toComma) + [IsIso i] : IsIso (homFromCommaOfIsIso i) := by + constructor + use homFromCommaOfIsIso (inv i) + constructor <;> ext : 1 <;> simp + +/-- Any isomorphism between objects of `P.Comma L R Q W` in `Comma L R` is also an isomorphism +in `P.Comma L R Q W`. -/ +@[simps] +def isoFromComma [Q.RespectsIso] [W.RespectsIso] {X Y : P.Comma L R Q W} + (i : X.toComma ≅ Y.toComma) : X ≅ Y where + hom := homFromCommaOfIsIso i.hom + inv := homFromCommaOfIsIso i.inv + +/-- Constructor for isomorphisms in `P.Comma L R Q W` from isomorphisms of the left and right +components and naturality in the forward direction. -/ +@[simps!] +def isoMk [Q.RespectsIso] [W.RespectsIso] {X Y : P.Comma L R Q W} (l : X.left ≅ Y.left) + (r : X.right ≅ Y.right) (h : L.map l.hom ≫ Y.hom = X.hom ≫ R.map r.hom := by aesop_cat) : + X ≅ Y := + isoFromComma (CategoryTheory.Comma.isoMk l r h) + +variable (L R P Q W) + +/-- The forgetful functor. -/ +@[simps] +def forget : P.Comma L R Q W ⥤ Comma L R where + obj X := X.toComma + map f := f.hom + +instance : (forget L R P Q W).Faithful where + map_injective := Comma.Hom.ext' + +variable {L R P Q W} + +instance {X Y : P.Comma L R Q W} (f : X ⟶ Y) [IsIso f] : IsIso f.hom := + (forget L R P Q W).map_isIso f + +lemma hom_homFromCommaOfIsIso [Q.RespectsIso] [W.RespectsIso] {X Y : P.Comma L R Q W} + (i : X ⟶ Y) [IsIso i.hom] : + homFromCommaOfIsIso i.hom = i := + rfl + +lemma inv_hom {X Y : P.Comma L R Q W} (f : X ⟶ Y) [IsIso f] : (inv f).hom = inv f.hom := by + apply IsIso.eq_inv_of_hom_inv_id + rw [← comp_hom, IsIso.hom_inv_id, id_hom] + +variable (L R P Q W) + +instance [Q.RespectsIso] [W.RespectsIso] : (forget L R P Q W).ReflectsIsomorphisms where + reflects f hf := by + simp only [forget_obj, forget_map] at hf + rw [← hom_homFromCommaOfIsIso f] + infer_instance + +/-- The forgetful functor from the full subcategory of `Comma L R` defined by `P` is +fully faithful. -/ +def forgetFullyFaithful : (forget L R P ⊤ ⊤).FullyFaithful where + preimage {X Y} f := ⟨f, trivial, trivial⟩ + +instance : (forget L R P ⊤ ⊤).Full := + Functor.FullyFaithful.full (forgetFullyFaithful L R P) + +end Comma + +end Comma + +section Over + +variable {T : Type*} [Category T] + +/-- Given a morphism property `P` on a category `C` and an object `X : C`, this is the +subcategory of `Over X` defined by `P` where morphisms satisfy `Q`. -/ +protected abbrev Over (P Q : MorphismProperty T) (X : T) : Type _ := + P.Comma (Functor.id T) (Functor.fromPUnit X) Q ⊤ + +variable (P Q : MorphismProperty T) [Q.IsMultiplicative] (X : T) + +/-- Construct a morphism in `P.Over Q X` from a morphism in `Over X`. -/ +@[simps hom] +def Over.Hom.mk {A B : P.Over Q X} (f : A.toComma ⟶ B.toComma) (hf : Q f.left) : A ⟶ B where + __ := f + prop_hom_left := hf + prop_hom_right := trivial + +/-- The forgetful functor from the full subcategory of `Over X` defined by `P` to `Over X`. -/ +protected abbrev Over.forget : P.Over Q X ⥤ Over X := + Comma.forget (Functor.id T) (Functor.fromPUnit X) P Q ⊤ + +end Over + +section Under + +variable {T : Type*} [Category T] + +/-- Given a morphism property `P` on a category `C` and an object `X : C`, this is the +subcategory of `Under X` defined by `P` where morphisms satisfy `Q`. -/ +protected abbrev Under (P Q : MorphismProperty T) (X : T) : Type _ := + P.Comma (Functor.fromPUnit X) (Functor.id T) ⊤ Q + +variable (P Q : MorphismProperty T) [Q.IsMultiplicative] (X : T) + +/-- Construct a morphism in `P.Under Q X` from a morphism in `Under.X`. -/ +@[simps hom] +def Under.Hom.mk {A B : P.Under Q X} (f : A.toComma ⟶ B.toComma) (hf : Q f.right) : A ⟶ B where + __ := f + prop_hom_left := trivial + prop_hom_right := hf + +/-- The forgetful functor from the full subcategory of `Under X` defined by `P` to `Under X`. -/ +protected abbrev Under.forget : P.Under Q X ⥤ Under X := + Comma.forget (Functor.fromPUnit X) (Functor.id T) P ⊤ Q + +end Under + +end CategoryTheory.MorphismProperty diff --git a/Mathlib/CategoryTheory/MorphismProperty/Composition.lean b/Mathlib/CategoryTheory/MorphismProperty/Composition.lean index 13fd8d28c5d30..9ceb89d7018f0 100644 --- a/Mathlib/CategoryTheory/MorphismProperty/Composition.lean +++ b/Mathlib/CategoryTheory/MorphismProperty/Composition.lean @@ -163,6 +163,10 @@ lemma of_op (W : MorphismProperty C) [IsMultiplicative W.op] : IsMultiplicative lemma of_unop (W : MorphismProperty Cᵒᵖ) [IsMultiplicative W.unop] : IsMultiplicative W := (inferInstance : IsMultiplicative W.unop.op) +instance : MorphismProperty.IsMultiplicative (⊤ : MorphismProperty C) where + comp_mem _ _ _ _ := trivial + id_mem _ := trivial + instance : (isomorphisms C).IsMultiplicative where id_mem _ := isomorphisms.infer_property _ comp_mem f g hf hg := by From 923081a67237679870311ae7e5efe20226bd28f4 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Thu, 24 Oct 2024 05:08:04 +0000 Subject: [PATCH 358/505] feat(AlgebraicGeometry/Restrict): more API for `Scheme.Opens`. (#17968) Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- Mathlib/AlgebraicGeometry/AffineScheme.lean | 10 +- .../Morphisms/Constructors.lean | 8 +- Mathlib/AlgebraicGeometry/Restrict.lean | 269 ++++++++++++------ 3 files changed, 185 insertions(+), 102 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/AffineScheme.lean b/Mathlib/AlgebraicGeometry/AffineScheme.lean index 35e6996d17928..8bf933f5421ce 100644 --- a/Mathlib/AlgebraicGeometry/AffineScheme.lean +++ b/Mathlib/AlgebraicGeometry/AffineScheme.lean @@ -342,12 +342,12 @@ theorem opensRange_fromSpec :hU.fromSpec.opensRange = U := Opens.ext (range_from @[reassoc (attr := simp)] theorem map_fromSpec {V : X.Opens} (hV : IsAffineOpen V) (f : op U ⟶ op V) : Spec.map (X.presheaf.map f) ≫ hU.fromSpec = hV.fromSpec := by - have : IsAffine (X.restrictFunctor.obj U).left := hU + have : IsAffine U := hU haveI : IsAffine _ := hV conv_rhs => - rw [fromSpec, ← X.restrictFunctor_map_ofRestrict f.unop, isoSpec_inv, Category.assoc, + rw [fromSpec, ← X.homOfLE_ι (V := U) f.unop.le, isoSpec_inv, Category.assoc, ← Scheme.isoSpec_inv_naturality_assoc, - ← Spec.map_comp_assoc, Scheme.restrictFunctor_map_app, ← Functor.map_comp] + ← Spec.map_comp_assoc, Scheme.homOfLE_app, ← Functor.map_comp] rw [fromSpec, isoSpec_inv, Category.assoc, ← Spec.map_comp_assoc, ← Functor.map_comp] rfl @@ -357,11 +357,11 @@ lemma Spec_map_appLE_fromSpec (f : X ⟶ Y) {V : X.Opens} {U : Y.Opens} Spec.map (f.appLE U V i) ≫ hU.fromSpec = hV.fromSpec ≫ f := by have : IsAffine U := hU simp only [IsAffineOpen.fromSpec, Category.assoc, isoSpec_inv] - simp_rw [← Scheme.restrictFunctor_map_ofRestrict (homOfLE i)] + simp_rw [← Scheme.homOfLE_ι _ i] rw [Category.assoc, ← morphismRestrict_ι, ← Category.assoc _ (f ∣_ U) U.ι, ← @Scheme.isoSpec_inv_naturality_assoc, ← Spec.map_comp_assoc, ← Spec.map_comp_assoc, Scheme.comp_app, morphismRestrict_app, - Scheme.restrictFunctor_map_app, Scheme.Hom.app_eq_appLE, Scheme.Hom.appLE_map, + Scheme.homOfLE_app, Scheme.Hom.app_eq_appLE, Scheme.Hom.appLE_map, Scheme.Hom.appLE_map, Scheme.Hom.appLE_map, Scheme.Hom.map_appLE] lemma fromSpec_top [IsAffine X] : (isAffineOpen_top X).fromSpec = X.isoSpec.inv := by diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Constructors.lean b/Mathlib/AlgebraicGeometry/Morphisms/Constructors.lean index bb4b68e99832c..1f5c0124ef2b8 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Constructors.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Constructors.lean @@ -167,15 +167,13 @@ theorem universally_isLocalAtTarget (P : MorphismProperty Scheme) rintro x - simpa using @hU (i₂.base x) trivial · rintro i - refine H _ ((X'.restrictIsoOfEq ?_).hom ≫ i₁ ∣_ _) (i₂ ∣_ _) _ ?_ + refine H _ ((X'.isoOfEq ?_).hom ≫ i₁ ∣_ _) (i₂ ∣_ _) _ ?_ · exact congr($(h.1.1) ⁻¹ᵁ U i) · rw [← (isPullback_morphismRestrict f _).paste_vert_iff] - · simp only [Scheme.restrictIsoOfEq, Category.assoc, morphismRestrict_ι, - IsOpenImmersion.isoOfRangeEq_hom_fac_assoc] + · simp only [Category.assoc, morphismRestrict_ι, Scheme.isoOfEq_hom_ι_assoc] exact (isPullback_morphismRestrict f' (i₂ ⁻¹ᵁ U i)).paste_vert h · rw [← cancel_mono (Scheme.Opens.ι _)] - simp [IsOpenImmersion.isoOfRangeEq_hom_fac_assoc, Scheme.restrictIsoOfEq, - morphismRestrict_ι_assoc, h.1.1] + simp [morphismRestrict_ι_assoc, h.1.1] end Universally diff --git a/Mathlib/AlgebraicGeometry/Restrict.lean b/Mathlib/AlgebraicGeometry/Restrict.lean index 2e8d4801e0e78..d58abb51d96f0 100644 --- a/Mathlib/AlgebraicGeometry/Restrict.lean +++ b/Mathlib/AlgebraicGeometry/Restrict.lean @@ -166,8 +166,8 @@ instance ΓRestrictAlgebra {X : Scheme.{u}} (U : X.Opens) : Algebra (Γ(X, ⊤)) Γ(U, ⊤) := (U.ι.app ⊤).toAlgebra -lemma Scheme.map_basicOpen' (r : Γ(U, ⊤)) : - U.ι ''ᵁ (U.toScheme.basicOpen r) = X.basicOpen +lemma Scheme.map_basicOpen (r : Γ(U, ⊤)) : + U.ι ''ᵁ U.toScheme.basicOpen r = X.basicOpen (X.presheaf.map (eqToHom U.isOpenEmbedding_obj_top.symm).op r) := by refine (Scheme.image_basicOpen (X.ofRestrict U.isOpenEmbedding) r).trans ?_ rw [← Scheme.basicOpen_res_eq _ _ (eqToHom U.isOpenEmbedding_obj_top).op] @@ -176,36 +176,85 @@ lemma Scheme.map_basicOpen' (r : Γ(U, ⊤)) : congr exact PresheafedSpace.IsOpenImmersion.ofRestrict_invApp _ _ _ +@[deprecated (since := "2024-10-23")] alias Scheme.map_basicOpen' := Scheme.map_basicOpen + lemma Scheme.Opens.ι_image_basicOpen (r : Γ(U, ⊤)) : - U.ι ''ᵁ (U.toScheme.basicOpen r) = X.basicOpen r := by - rw [Scheme.map_basicOpen', Scheme.basicOpen_res_eq] + U.ι ''ᵁ U.toScheme.basicOpen r = X.basicOpen r := by + rw [Scheme.map_basicOpen, Scheme.basicOpen_res_eq] lemma Scheme.map_basicOpen_map (r : Γ(X, U)) : U.ι ''ᵁ (U.toScheme.basicOpen <| U.topIso.inv r) = X.basicOpen r := by simp only [Scheme.Opens.toScheme_presheaf_obj] - rw [Scheme.map_basicOpen', Scheme.basicOpen_res_eq, Scheme.Opens.topIso_inv, + rw [Scheme.map_basicOpen, Scheme.basicOpen_res_eq, Scheme.Opens.topIso_inv, Scheme.basicOpen_res_eq X] +/-- If `U ≤ V`, then `U` is also a subscheme of `V`. -/ +protected noncomputable +def Scheme.homOfLE (X : Scheme.{u}) {U V : X.Opens} (e : U ≤ V) : (U : Scheme.{u}) ⟶ V := + IsOpenImmersion.lift V.ι U.ι (by simpa using e) + +@[reassoc (attr := simp)] +lemma Scheme.homOfLE_ι (X : Scheme.{u}) {U V : X.Opens} (e : U ≤ V) : + X.homOfLE e ≫ V.ι = U.ι := + IsOpenImmersion.lift_fac _ _ _ + +@[simp] +lemma Scheme.homOfLE_rfl (X : Scheme.{u}) (U : X.Opens) : X.homOfLE (refl U) = 𝟙 _ := by + rw [← cancel_mono U.ι, Scheme.homOfLE_ι, Category.id_comp] + +@[reassoc (attr := simp)] +lemma Scheme.homOfLE_homOfLE (X : Scheme.{u}) {U V W : X.Opens} (e₁ : U ≤ V) (e₂ : V ≤ W) : + X.homOfLE e₁ ≫ X.homOfLE e₂ = X.homOfLE (e₁.trans e₂) := by + rw [← cancel_mono W.ι, Category.assoc, Scheme.homOfLE_ι, Scheme.homOfLE_ι, Scheme.homOfLE_ι] + +theorem Scheme.homOfLE_base {U V : X.Opens} (e : U ≤ V) : + (X.homOfLE e).base = (Opens.toTopCat _).map (homOfLE e) := by + ext a; refine Subtype.ext ?_ -- Porting note: `ext` did not pick up `Subtype.ext` + exact congr($(X.homOfLE_ι e).base a) + +@[simp] +theorem Scheme.homOfLE_apply {U V : X.Opens} (e : U ≤ V) (x : U) : + ((X.homOfLE e).base x).1 = x := by + rw [homOfLE_base] + rfl + +theorem Scheme.ι_image_homOfLE_le_ι_image {U V : X.Opens} (e : U ≤ V) (W : Opens V) : + U.ι ''ᵁ (X.homOfLE e ⁻¹ᵁ W) ≤ V.ι ''ᵁ W := by + simp only [← SetLike.coe_subset_coe, IsOpenMap.functor_obj_coe, Set.image_subset_iff, + Scheme.homOfLE_base, Opens.map_coe, Opens.inclusion'_apply] + rintro _ h + exact ⟨_, h, rfl⟩ + +@[simp] +theorem Scheme.homOfLE_app {U V : X.Opens} (e : U ≤ V) (W : Opens V) : + (X.homOfLE e).app W = + X.presheaf.map (homOfLE <| X.ι_image_homOfLE_le_ι_image e W).op := by + have e₁ := Scheme.congr_app (X.homOfLE_ι e) (V.ι ''ᵁ W) + have : V.ι ⁻¹ᵁ V.ι ''ᵁ W = W := W.map_functor_eq (U := V) + have e₂ := (X.homOfLE e).naturality (eqToIso this).hom.op + have e₃ := e₂.symm.trans e₁ + dsimp at e₃ ⊢ + rw [← IsIso.eq_comp_inv, ← Functor.map_inv, ← Functor.map_comp] at e₃ + rw [e₃, ← Functor.map_comp] + congr 1 + +instance (X : Scheme.{u}) {U V : X.Opens} (e : U ≤ V) : IsOpenImmersion (X.homOfLE e) := by + delta Scheme.homOfLE + infer_instance + -- Porting note: `simps` can't synthesize `obj_left, obj_hom, mapLeft` variable (X) in /-- The functor taking open subsets of `X` to open subschemes of `X`. -/ -- @[simps obj_left obj_hom mapLeft] def Scheme.restrictFunctor : X.Opens ⥤ Over X where obj U := Over.mk U.ι - map {U V} i := - Over.homMk - (IsOpenImmersion.lift V.ι U.ι <| by simpa using i.le) - (IsOpenImmersion.lift_fac _ _ _) + map {U V} i := Over.homMk (X.homOfLE i.le) (by simp) map_id U := by ext1 - dsimp only [Over.homMk_left, Over.id_left] - rw [← cancel_mono U.ι, Category.id_comp, - IsOpenImmersion.lift_fac] + exact Scheme.homOfLE_rfl _ _ map_comp {U V W} i j := by ext1 - dsimp only [Over.homMk_left, Over.comp_left] - rw [← cancel_mono W.ι, Category.assoc] - iterate 3 rw [IsOpenImmersion.lift_fac] + exact (X.homOfLE_homOfLE i.le j.le).symm @[simp] lemma Scheme.restrictFunctor_obj_left (U : X.Opens) : (X.restrictFunctor.obj U).left = U := rfl @@ -213,45 +262,21 @@ def Scheme.restrictFunctor : X.Opens ⥤ Over X where @[simp] lemma Scheme.restrictFunctor_obj_hom (U : X.Opens) : (X.restrictFunctor.obj U).hom = U.ι := rfl -/-- This is not a `simp` lemma, as `(X.restricFunctor.map i).left` is used as the `simp` -normal-form for the induced morphism `U.toScheme ⟶ V.toScheme`. -/ +@[simp] lemma Scheme.restrictFunctor_map_left {U V : X.Opens} (i : U ⟶ V) : - (X.restrictFunctor.map i).left = IsOpenImmersion.lift (V.ι) U.ι (by simpa using i.le) := rfl + (X.restrictFunctor.map i).left = (X.homOfLE i.le) := rfl -instance {U V : X.Opens} (i : U ⟶ V) : IsOpenImmersion (X.restrictFunctor.map i).left := by - rw [X.restrictFunctor_map_left] - infer_instance +@[deprecated (since := "2024-10-20")] +alias Scheme.restrictFunctor_map_ofRestrict := Scheme.homOfLE_ι +@[deprecated (since := "2024-10-20")] +alias Scheme.restrictFunctor_map_ofRestrict_assoc := Scheme.homOfLE_ι_assoc --- Porting note: the `by ...` used to be automatically done by unification magic -@[reassoc] -theorem Scheme.restrictFunctor_map_ofRestrict {U V : X.Opens} (i : U ⟶ V) : - (X.restrictFunctor.map i).left ≫ V.ι = U.ι := - IsOpenImmersion.lift_fac _ _ (by simpa using i.le) - -theorem Scheme.restrictFunctor_map_base {U V : X.Opens} (i : U ⟶ V) : - (X.restrictFunctor.map i).left.base = (Opens.toTopCat _).map i := by - ext a; refine Subtype.ext ?_ -- Porting note: `ext` did not pick up `Subtype.ext` - exact (congr_arg (fun f : X.restrict U.isOpenEmbedding ⟶ X => f.base a) - (X.restrictFunctor_map_ofRestrict i)) - -theorem Scheme.restrictFunctor_map_app_aux {U V : X.Opens} (i : U ⟶ V) (W : Opens V) : - U.ι ''ᵁ ((X.restrictFunctor.map i).left ⁻¹ᵁ W) ≤ V.ι ''ᵁ W := by - simp only [← SetLike.coe_subset_coe, IsOpenMap.functor_obj_coe, Set.image_subset_iff, - Scheme.restrictFunctor_map_base, Opens.map_coe, Opens.inclusion'_apply] - rintro _ h - exact ⟨_, h, rfl⟩ - -theorem Scheme.restrictFunctor_map_app {U V : X.Opens} (i : U ⟶ V) (W : Opens V) : - (X.restrictFunctor.map i).1.app W = - X.presheaf.map (homOfLE <| X.restrictFunctor_map_app_aux i W).op := by - have e₁ := Scheme.congr_app (X.restrictFunctor_map_ofRestrict i) (V.ι ''ᵁ W) - have : V.ι ⁻¹ᵁ V.ι ''ᵁ W = W := W.map_functor_eq (U := V) - have e₂ := (X.restrictFunctor.map i).1.naturality (eqToIso this).hom.op - have e₃ := e₂.symm.trans e₁ - dsimp at e₃ ⊢ - rw [← IsIso.eq_comp_inv, ← Functor.map_inv, ← Functor.map_comp] at e₃ - rw [e₃, ← Functor.map_comp] - congr 1 +@[deprecated (since := "2024-10-20")] +alias Scheme.restrictFunctor_map_base := Scheme.homOfLE_base +@[deprecated (since := "2024-10-20")] +alias Scheme.restrictFunctor_map_app_aux := Scheme.ι_image_homOfLE_le_ι_image +@[deprecated (since := "2024-10-20")] +alias Scheme.restrictFunctor_map_app := Scheme.homOfLE_app /-- The functor that restricts to open subschemes and then takes global section is isomorphic to the structure sheaf. -/ @@ -262,7 +287,7 @@ def Scheme.restrictFunctorΓ : X.restrictFunctor.op ⋙ (Over.forget X).op ⋙ S (by intro U V i dsimp - rw [X.restrictFunctor_map_app, ← Functor.map_comp, ← Functor.map_comp] + rw [X.homOfLE_app, ← Functor.map_comp, ← Functor.map_comp] congr 1) /-- `X ∣_ U ∣_ V` is isomorphic to `X ∣_ V ∣_ U` -/ @@ -272,40 +297,98 @@ def Scheme.restrictRestrictComm (X : Scheme.{u}) (U V : X.Opens) : IsOpenImmersion.isoOfRangeEq (Opens.ι _ ≫ U.ι) (Opens.ι _ ≫ V.ι) <| by simp [Set.image_preimage_eq_inter_range, Set.inter_comm (U : Set X), Set.range_comp] -/-- If `V` is an open subset of `U`, then `X ∣_ U ∣_ V` is isomorphic to `X ∣_ V`. -/ +/-- If `f : X ⟶ Y` is an open immersion, then for any `U : X.Opens`, +we have the isomorphism `U ≅ f ''ᵁ U`. -/ noncomputable -def Scheme.restrictRestrict (X : Scheme.{u}) (U : X.Opens) (V : U.toScheme.Opens) : - V.toScheme ≅ U.ι ''ᵁ V := - IsOpenImmersion.isoOfRangeEq (Opens.ι _ ≫ U.ι) (Opens.ι _) (by simp [Set.range_comp]) +def Scheme.Hom.isoImage + {X Y : Scheme.{u}} (f : X.Hom Y) [IsOpenImmersion f] (U : X.Opens) : + U.toScheme ≅ f ''ᵁ U := + IsOpenImmersion.isoOfRangeEq (Opens.ι _ ≫ f) (Opens.ι _) (by simp [Set.range_comp]) -@[simp, reassoc] -lemma Scheme.restrictRestrict_hom_restrict (X : Scheme.{u}) (U : X.Opens) - (V : U.toScheme.Opens) : - (X.restrictRestrict U V).hom ≫ Opens.ι _ = V.ι ≫ U.ι := +@[reassoc (attr := simp)] +lemma Scheme.Hom.isoImage_hom_ι + {X Y : Scheme.{u}} (f : X ⟶ Y) [IsOpenImmersion f] (U : X.Opens) : + (f.isoImage U).hom ≫ (f ''ᵁ U).ι = U.ι ≫ f := IsOpenImmersion.isoOfRangeEq_hom_fac _ _ _ -@[simp, reassoc] -lemma Scheme.restrictRestrict_inv_restrict_restrict (X : Scheme.{u}) (U : X.Opens) - (V : U.toScheme.Opens) : - (X.restrictRestrict U V).inv ≫ V.ι ≫ U.ι = Opens.ι _ := +@[reassoc (attr := simp)] +lemma Scheme.Hom.isoImage_inv_ι + {X Y : Scheme.{u}} (f : X ⟶ Y) [IsOpenImmersion f] (U : X.Opens) : + (f.isoImage U).inv ≫ U.ι ≫ f = (f ''ᵁ U).ι := IsOpenImmersion.isoOfRangeEq_inv_fac _ _ _ +@[deprecated (since := "2024-10-20")] +alias Scheme.restrictRestrict := Scheme.Hom.isoImage +@[deprecated (since := "2024-10-20")] +alias Scheme.restrictRestrict_hom_restrict := Scheme.Hom.isoImage_hom_ι +@[deprecated (since := "2024-10-20")] +alias Scheme.restrictRestrict_inv_restrict_restrict := Scheme.Hom.isoImage_inv_ι +@[deprecated (since := "2024-10-20")] +alias Scheme.restrictRestrict_hom_restrict_assoc := Scheme.Hom.isoImage_hom_ι_assoc +@[deprecated (since := "2024-10-20")] +alias Scheme.restrictRestrict_inv_restrict_restrict_assoc := Scheme.Hom.isoImage_inv_ι_assoc + +/-- `(⊤ : X.Opens)` as a scheme is isomorphic to `X`. -/ +@[simps hom] +def Scheme.topIso (X : Scheme) : ↑(⊤ : X.Opens) ≅ X where + hom := Scheme.Opens.ι _ + inv := ⟨X.restrictTopIso.inv⟩ + hom_inv_id := Hom.ext' X.restrictTopIso.hom_inv_id + inv_hom_id := Hom.ext' X.restrictTopIso.inv_hom_id + +@[reassoc (attr := simp)] +lemma Scheme.toIso_inv_ι (X : Scheme.{u}) : X.topIso.inv ≫ Opens.ι _ = 𝟙 _ := + X.topIso.inv_hom_id + +@[reassoc (attr := simp)] +lemma Scheme.ι_toIso_inv (X : Scheme.{u}) : Opens.ι _ ≫ X.topIso.inv = 𝟙 _ := + X.topIso.hom_inv_id + /-- If `U = V`, then `X ∣_ U` is isomorphic to `X ∣_ V`. -/ noncomputable -def Scheme.restrictIsoOfEq (X : Scheme.{u}) {U V : X.Opens} (e : U = V) : +def Scheme.isoOfEq (X : Scheme.{u}) {U V : X.Opens} (e : U = V) : (U : Scheme.{u}) ≅ V := - IsOpenImmersion.isoOfRangeEq U.ι (V.ι) (by rw [e]) + IsOpenImmersion.isoOfRangeEq U.ι V.ι (by rw [e]) + +@[reassoc (attr := simp)] +lemma Scheme.isoOfEq_hom_ι (X : Scheme.{u}) {U V : X.Opens} (e : U = V) : + (X.isoOfEq e).hom ≫ V.ι = U.ι := + IsOpenImmersion.isoOfRangeEq_hom_fac _ _ _ + +@[reassoc (attr := simp)] +lemma Scheme.isoOfEq_inv_ι (X : Scheme.{u}) {U V : X.Opens} (e : U = V) : + (X.isoOfEq e).inv ≫ U.ι = V.ι := + IsOpenImmersion.isoOfRangeEq_inv_fac _ _ _ + +@[simp] +lemma Scheme.isoOfEq_rfl (X : Scheme.{u}) (U : X.Opens) : X.isoOfEq (refl U) = Iso.refl _ := by + ext1 + rw [← cancel_mono U.ι, Scheme.isoOfEq_hom_ι, Iso.refl_hom, Category.id_comp] + +@[deprecated (since := "2024-10-20")] alias Scheme.restrictIsoOfEq := Scheme.isoOfEq end /-- The restriction of an isomorphism onto an open set. -/ -noncomputable abbrev Scheme.restrictMapIso {X Y : Scheme.{u}} (f : X ⟶ Y) [IsIso f] +noncomputable def Scheme.Hom.preimageIso {X Y : Scheme.{u}} (f : X.Hom Y) [IsIso (C := Scheme) f] (U : Y.Opens) : (f ⁻¹ᵁ U).toScheme ≅ U := by apply IsOpenImmersion.isoOfRangeEq (f := (f ⁻¹ᵁ U).ι ≫ f) U.ι _ dsimp rw [Set.range_comp, Opens.range_ι, Opens.range_ι] refine @Set.image_preimage_eq _ _ f.base U.1 f.homeomorph.surjective +@[reassoc (attr := simp)] +lemma Scheme.Hom.preimageIso_hom_ι {X Y : Scheme.{u}} (f : X.Hom Y) [IsIso (C := Scheme) f] + (U : Y.Opens) : (f.preimageIso U).hom ≫ U.ι = (f ⁻¹ᵁ U).ι ≫ f := + IsOpenImmersion.isoOfRangeEq_hom_fac _ _ _ + +@[reassoc (attr := simp)] +lemma Scheme.Hom.preimageIso_inv_ι {X Y : Scheme.{u}} (f : X.Hom Y) [IsIso (C := Scheme) f] + (U : Y.Opens) : (f.preimageIso U).inv ≫ (f ⁻¹ᵁ U).ι ≫ f = U.ι := + IsOpenImmersion.isoOfRangeEq_inv_fac _ _ _ + +@[deprecated (since := "2024-10-20")] alias Scheme.restrictMapIso := Scheme.Hom.preimageIso + section MorphismRestrict /-- Given a morphism `f : X ⟶ Y` and an open set `U ⊆ Y`, we have `X ×[Y] U ≅ X |_{f ⁻¹ U}` -/ @@ -319,11 +402,16 @@ theorem pullbackRestrictIsoRestrict_inv_fst {X Y : Scheme.{u}} (f : X ⟶ Y) (U (pullbackRestrictIsoRestrict f U).inv ≫ pullback.fst f _ = (f ⁻¹ᵁ U).ι := by delta pullbackRestrictIsoRestrict; simp -@[simp, reassoc] -theorem pullbackRestrictIsoRestrict_hom_restrict {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) : +@[reassoc (attr := simp)] +theorem pullbackRestrictIsoRestrict_hom_ι {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) : (pullbackRestrictIsoRestrict f U).hom ≫ (f ⁻¹ᵁ U).ι = pullback.fst f _ := by delta pullbackRestrictIsoRestrict; simp +@[deprecated (since := "2024-10-20")] +alias pullbackRestrictIsoRestrict_hom_restrict := pullbackRestrictIsoRestrict_hom_ι +@[deprecated (since := "2024-10-20")] +alias pullbackRestrictIsoRestrict_hom_restrict_assoc := pullbackRestrictIsoRestrict_hom_ι_assoc + /-- The restriction of a morphism `X ⟶ Y` onto `X |_{f ⁻¹ U} ⟶ Y |_ U`. -/ def morphismRestrict {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) : (f ⁻¹ᵁ U).toScheme ⟶ U := (pullbackRestrictIsoRestrict f U).inv ≫ pullback.snd _ _ @@ -331,12 +419,12 @@ def morphismRestrict {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) : (f ⁻¹ /-- the notation for restricting a morphism of scheme to an open subset of the target scheme -/ infixl:85 " ∣_ " => morphismRestrict -@[simp, reassoc] +@[reassoc (attr := simp)] theorem pullbackRestrictIsoRestrict_hom_morphismRestrict {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) : (pullbackRestrictIsoRestrict f U).hom ≫ f ∣_ U = pullback.snd _ _ := Iso.hom_inv_id_assoc _ _ -@[simp, reassoc] +@[reassoc (attr := simp)] theorem morphismRestrict_ι {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) : (f ∣_ U) ≫ U.ι = (f ⁻¹ᵁ U).ι ≫ f := by delta morphismRestrict @@ -455,7 +543,7 @@ def morphismRestrictOpensRange rw [Iso.trans_hom, asIso_hom, ← Iso.comp_inv_eq, ← cancel_mono g, Arrow.mk_hom, Arrow.mk_hom, Category.assoc, Category.assoc, Category.assoc, IsOpenImmersion.isoOfRangeEq_inv_fac, ← pullback.condition, morphismRestrict_ι, - pullbackRestrictIsoRestrict_hom_restrict_assoc, pullback.lift_fst_assoc, Category.comp_id] + pullbackRestrictIsoRestrict_hom_ι_assoc, pullback.lift_fst_assoc, Category.comp_id] /-- The restrictions onto two equal open sets are isomorphic. This currently has bad defeqs when unfolded, but it should not matter for now. Replace this definition if better defeqs are needed. -/ @@ -466,8 +554,8 @@ def morphismRestrictEq {X Y : Scheme.{u}} (f : X ⟶ Y) {U V : Y.Opens} (e : U = /-- Restricting a morphism twice is isomorphic to one restriction. -/ def morphismRestrictRestrict {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) (V : U.toScheme.Opens) : Arrow.mk (f ∣_ U ∣_ V) ≅ Arrow.mk (f ∣_ U.ι ''ᵁ V) := by - refine Arrow.isoMk' _ _ (Scheme.restrictRestrict _ _ _ ≪≫ Scheme.restrictIsoOfEq _ ?_) - (Scheme.restrictRestrict _ _ _) ?_ + refine Arrow.isoMk' _ _ ((Scheme.Opens.ι _).isoImage _ ≪≫ Scheme.isoOfEq _ ?_) + ((Scheme.Opens.ι _).isoImage _) ?_ · ext x simp only [IsOpenMap.functor_obj_coe, Opens.coe_inclusion', Opens.map_coe, Set.mem_image, Set.mem_preimage, SetLike.mem_coe, morphismRestrict_base] @@ -477,10 +565,9 @@ def morphismRestrictRestrict {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) (V : · rintro ⟨⟨a, h₁⟩, h₂, rfl : a = _⟩ exact ⟨⟨x, h₁⟩, h₂, rfl⟩ · rw [← cancel_mono (Scheme.Opens.ι _), Iso.trans_hom, Category.assoc, Category.assoc, - Category.assoc, morphismRestrict_ι, Scheme.restrictIsoOfEq, - IsOpenImmersion.isoOfRangeEq_hom_fac_assoc, - Scheme.restrictRestrict_hom_restrict_assoc, - Scheme.restrictRestrict_hom_restrict, + Category.assoc, morphismRestrict_ι, Scheme.isoOfEq_hom_ι_assoc, + Scheme.Hom.isoImage_hom_ι_assoc, + Scheme.Hom.isoImage_hom_ι, morphismRestrict_ι_assoc, morphismRestrict_ι] /-- Restricting a morphism twice onto a basic open set is isomorphic to one restriction. -/ @@ -524,20 +611,20 @@ namespace Scheme.Hom /-- The restriction of a morphism `f : X ⟶ Y` to open sets on the source and target. -/ def resLE (f : Hom X Y) (U : Y.Opens) (V : X.Opens) (e : V ≤ f ⁻¹ᵁ U) : V.toScheme ⟶ U.toScheme := - (X.restrictFunctor.map (homOfLE e)).left ≫ f ∣_ U + X.homOfLE e ≫ f ∣_ U variable (f : X ⟶ Y) {U U' : Y.Opens} {V V' : X.Opens} (e : V ≤ f ⁻¹ᵁ U) lemma resLE_eq_morphismRestrict : f.resLE U (f ⁻¹ᵁ U) le_rfl = f ∣_ U := by - simp [Scheme.Hom.resLE] + simp [resLE] -lemma resLE_id (i : V ⟶ V') : resLE (𝟙 X) V' V i.le = (X.restrictFunctor.map i).left := by +lemma resLE_id (i : V ≤ V') : resLE (𝟙 X) V' V i = X.homOfLE i := by simp only [resLE, morphismRestrict_id] rfl @[reassoc (attr := simp)] lemma resLE_comp_ι : f.resLE U V e ≫ U.ι = V.ι ≫ f := by - simp [resLE, restrictFunctor_map_ofRestrict_assoc] + simp [resLE] @[reassoc] lemma resLE_comp_resLE {Z : Scheme.{u}} (g : Y ⟶ Z) {W : Z.Opens} (e') : @@ -546,14 +633,14 @@ lemma resLE_comp_resLE {Z : Scheme.{u}} (g : Y ⟶ Z) {W : Z.Opens} (e') : simp [← cancel_mono W.ι] @[reassoc (attr := simp)] -lemma map_resLE (i : V' ⟶ V) : - (X.restrictFunctor.map i).left ≫ f.resLE U V e = f.resLE U V' (i.le.trans e) := by +lemma map_resLE (i : V' ≤ V) : + X.homOfLE i ≫ f.resLE U V e = f.resLE U V' (i.trans e) := by simp_rw [← resLE_id, resLE_comp_resLE, Category.id_comp] @[reassoc (attr := simp)] -lemma resLE_map (i : U ⟶ U') : - f.resLE U V e ≫ (Y.restrictFunctor.map i).left = - f.resLE U' V (e.trans ((Opens.map f.base).map i).le) := by +lemma resLE_map (i : U ≤ U') : + f.resLE U V e ≫ Y.homOfLE i = + f.resLE U' V (e.trans ((Opens.map f.base).map i.hom).le) := by simp_rw [← resLE_id, resLE_comp_resLE, Category.comp_id] lemma resLE_congr (e₁ : U = U') (e₂ : V = V') (P : MorphismProperty Scheme.{u}) : @@ -574,12 +661,10 @@ lemma resLE_appLE {U : Y.Opens} {V : X.Opens} (e : V ≤ f ⁻¹ᵁ U) (O : U.toScheme.Opens) (W : V.toScheme.Opens) (e' : W ≤ resLE f U V e ⁻¹ᵁ O) : (f.resLE U V e).appLE O W e' = f.appLE (U.ι ''ᵁ O) (V.ι ''ᵁ W) ((le_preimage_resLE_iff f e O W).mp e') := by - simp only [Scheme.Hom.appLE, Scheme.Hom.resLE, Scheme.restrictFunctor_map_left, Opens.map_coe, - id_eq, Scheme.comp_app, morphismRestrict_app', Category.assoc, IsOpenImmersion.lift_app, - Scheme.Opens.ι_appIso, Scheme.Opens.ι_app, Scheme.Opens.toScheme_presheaf_map, Category.assoc] + simp only [appLE, resLE, comp_coeBase, Opens.map_comp_obj, comp_app, morphismRestrict_app', + homOfLE_leOfHom, homOfLE_app, Category.assoc, Opens.toScheme_presheaf_map, Quiver.Hom.unop_op, + opensFunctor_map_homOfLE] rw [← X.presheaf.map_comp, ← X.presheaf.map_comp] - erw [Category.id_comp] - rw [← X.presheaf.map_comp] rfl end Scheme.Hom From cd1f49389b69f1e427def7a7c0ddc46f6a7da7dd Mon Sep 17 00:00:00 2001 From: mathlib4-update-dependencies-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Thu, 24 Oct 2024 05:46:18 +0000 Subject: [PATCH 359/505] chore: update Mathlib dependencies 2024-10-24 (#18155) This PR updates the Mathlib dependencies. --- lake-manifest.json | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lake-manifest.json b/lake-manifest.json index 02b5a77289c00..7a31bfbacb159 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -55,7 +55,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "9b4088ccf0f44ddd7b1132bb1348aef8cf481e12", + "rev": "0ea83a676d288220ba227808568cbb80fe43ace0", "name": "importGraph", "manifestFile": "lake-manifest.json", "inputRev": "main", From 893b3a846645d48ac45143ec02149cc551acd99d Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= <37772949+joelriou@users.noreply.github.com> Date: Thu, 24 Oct 2024 07:23:49 +0000 Subject: [PATCH 360/505] feat(Algebra/Category/ModuleCat): the free presheaf of modules on a presheaf of sets (#16755) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This contribution was created as part of the AIM workshop "Formalizing algebraic geometry" in June 2024. Co-authored-by: Johan Commelin Co-authored-by: Joël Riou Co-authored-by: Johan Commelin --- Mathlib.lean | 1 + .../Category/ModuleCat/Adjunctions.lean | 53 +++++++- .../Category/ModuleCat/Presheaf/Free.lean | 113 ++++++++++++++++++ .../ModuleCat/Presheaf/Sheafification.lean | 2 +- .../GroupCohomology/Resolution.lean | 3 +- 5 files changed, 163 insertions(+), 9 deletions(-) create mode 100644 Mathlib/Algebra/Category/ModuleCat/Presheaf/Free.lean diff --git a/Mathlib.lean b/Mathlib.lean index 488148a28f8e6..0511dec58327e 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -114,6 +114,7 @@ import Mathlib.Algebra.Category.ModuleCat.Presheaf import Mathlib.Algebra.Category.ModuleCat.Presheaf.Abelian import Mathlib.Algebra.Category.ModuleCat.Presheaf.ChangeOfRings import Mathlib.Algebra.Category.ModuleCat.Presheaf.Colimits +import Mathlib.Algebra.Category.ModuleCat.Presheaf.Free import Mathlib.Algebra.Category.ModuleCat.Presheaf.Limits import Mathlib.Algebra.Category.ModuleCat.Presheaf.Pushforward import Mathlib.Algebra.Category.ModuleCat.Presheaf.Sheafification diff --git a/Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean b/Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean index 7d7e0e34c73af..229bb1717e4b4 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean @@ -33,22 +33,63 @@ variable [Ring R] /-- The free functor `Type u ⥤ ModuleCat R` sending a type `X` to the free `R`-module with generators `x : X`, implemented as the type `X →₀ R`. -/ -@[simps] def free : Type u ⥤ ModuleCat R where obj X := ModuleCat.of R (X →₀ R) map {_ _} f := Finsupp.lmapDomain _ _ f map_id := by intros; exact Finsupp.lmapDomain_id _ _ map_comp := by intros; exact Finsupp.lmapDomain_comp _ _ _ _ +variable {R} + +/-- Constructor for elements in the module `(free R).obj X`. -/ +noncomputable def freeMk {X : Type u} (x : X) : (free R).obj X := Finsupp.single x 1 + +@[ext 1200] +lemma free_hom_ext {X : Type u} {M : ModuleCat.{u} R} {f g : (free R).obj X ⟶ M} + (h : ∀ (x : X), f (freeMk x) = g (freeMk x)) : + f = g := + (Finsupp.lhom_ext' (fun x ↦ LinearMap.ext_ring (h x))) + +/-- The morphism of modules `(free R).obj X ⟶ M` corresponding +to a map `f : X ⟶ M`. -/ +noncomputable def freeDesc {X : Type u} {M : ModuleCat.{u} R} (f : X ⟶ M) : + (free R).obj X ⟶ M := + Finsupp.lift M R X f + +@[simp] +lemma freeDesc_apply {X : Type u} {M : ModuleCat.{u} R} (f : X ⟶ M) (x : X) : + freeDesc f (freeMk x) = f x := by + dsimp [freeDesc] + erw [Finsupp.lift_apply, Finsupp.sum_single_index] + all_goals simp + +@[simp] +lemma free_map_apply {X Y : Type u} (f : X → Y) (x : X) : + (free R).map f (freeMk x) = freeMk (f x) := by + apply Finsupp.mapDomain_single + +/-- The bijection `((free R).obj X ⟶ M) ≃ (X → M)` when `X` is a type and `M` a module. -/ +@[simps] +def freeHomEquiv {X : Type u} {M : ModuleCat.{u} R} : + ((free R).obj X ⟶ M) ≃ (X → M) where + toFun φ x := φ (freeMk x) + invFun ψ := freeDesc ψ + left_inv _ := by ext; simp + right_inv _ := by ext; simp + +variable (R) + /-- The free-forgetful adjunction for R-modules. -/ def adj : free R ⊣ forget (ModuleCat.{u} R) := Adjunction.mkOfHomEquiv - { homEquiv := fun X M => (Finsupp.lift M R X).toEquiv.symm - homEquiv_naturality_left_symm := fun {_ _} M _ g => - Finsupp.lhom_ext' fun _ => - LinearMap.ext_ring - (Finsupp.sum_mapDomain_index_addMonoidHom fun y => (smulAddHom R M).flip (g y)).symm } + { homEquiv := fun _ _ => freeHomEquiv + homEquiv_naturality_left_symm := fun {X Y M} f g ↦ by ext; simp [freeHomEquiv] } + +@[simp] +lemma adj_homEquiv (X : Type u) (M : ModuleCat.{u} R) : + (adj R).homEquiv X M = freeHomEquiv := by + simp only [adj, Adjunction.mkOfHomEquiv_homEquiv] instance : (forget (ModuleCat.{u} R)).IsRightAdjoint := (adj R).isRightAdjoint diff --git a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Free.lean b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Free.lean new file mode 100644 index 0000000000000..e17dd748c6cfd --- /dev/null +++ b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Free.lean @@ -0,0 +1,113 @@ +/- +Copyright (c) 2024 Johan Commelin. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johan Commelin, Joel Riou +-/ +import Mathlib.Algebra.Category.ModuleCat.Presheaf +import Mathlib.Algebra.Category.ModuleCat.Adjunctions + +/-! +# The free presheaf of modules on a presheaf of sets + +In this file, given a presheaf of rings `R` on a category `C`, +we construct the functor +`PresheafOfModules.free : (Cᵒᵖ ⥤ Type u) ⥤ PresheafOfModules.{u} R` +which sends a presheaf of types to the corresponding presheaf of free modules. +`PresheafOfModules.freeAdjunction` shows that this functor is the left +adjoint to the forget functor. + +## Notes + +This contribution was created as part of the AIM workshop +"Formalizing algebraic geometry" in June 2024. + +-/ + +universe u v₁ u₁ + +open CategoryTheory + +namespace PresheafOfModules + +variable {C : Type u₁} [Category.{v₁} C] (R : Cᵒᵖ ⥤ RingCat.{u}) + +variable {R} in +/-- Given a presheaf of types `F : Cᵒᵖ ⥤ Type u`, this is the presheaf +of modules over `R` which sends `X : Cᵒᵖ` to the free `R.obj X`-module on `F.obj X`. -/ +@[simps] +noncomputable def freeObj (F : Cᵒᵖ ⥤ Type u) : PresheafOfModules.{u} R where + obj X := (ModuleCat.free (R.obj X)).obj (F.obj X) + map {X Y} f := ModuleCat.freeDesc (fun x ↦ ModuleCat.freeMk (F.map f x)) + map_id := by aesop + +/-- The free presheaf of modules functor `(Cᵒᵖ ⥤ Type u) ⥤ PresheafOfModules.{u} R`. -/ +@[simps] +noncomputable def free : (Cᵒᵖ ⥤ Type u) ⥤ PresheafOfModules.{u} R where + obj := freeObj + map {F G} φ := + { app := fun X ↦ (ModuleCat.free (R.obj X)).map (φ.app X) + naturality := fun {X Y} f ↦ by + dsimp + ext x + simp [FunctorToTypes.naturality] } + +section + +variable {R} + +variable {F : Cᵒᵖ ⥤ Type u} {G : PresheafOfModules.{u} R} + +/-- The morphism of presheaves of modules `freeObj F ⟶ G` corresponding to +a morphism `F ⟶ G.presheaf ⋙ forget _` of presheaves of types. -/ +@[simps] +noncomputable def freeObjDesc (φ : F ⟶ G.presheaf ⋙ forget _) : freeObj F ⟶ G where + app X := ModuleCat.freeDesc (φ.app X) + naturality {X Y} f := by + dsimp + ext x + simpa using NatTrans.naturality_apply φ f x + +variable (F R) in +/-- The unit of `PresheafOfModules.freeAdjunction`. -/ +@[simps] +noncomputable def freeAdjunctionUnit : F ⟶ (freeObj (R := R) F).presheaf ⋙ forget _ where + app X x := ModuleCat.freeMk x + naturality X Y f := by ext; simp [presheaf] + +/-- The bijection `(freeObj F ⟶ G) ≃ (F ⟶ G.presheaf ⋙ forget _)` when +`F` is a presheaf of types and `G` a presheaf of modules. -/ +noncomputable def freeHomEquiv : (freeObj F ⟶ G) ≃ (F ⟶ G.presheaf ⋙ forget _) where + toFun ψ := freeAdjunctionUnit R F ≫ whiskerRight ((toPresheaf _).map ψ) _ + invFun φ := freeObjDesc φ + left_inv ψ := by ext1 X; dsimp; ext x; simp [toPresheaf] + right_inv φ := by ext; simp [toPresheaf] + +lemma free_hom_ext {ψ ψ' : freeObj F ⟶ G} + (h : freeAdjunctionUnit R F ≫ whiskerRight ((toPresheaf _).map ψ) _ = + freeAdjunctionUnit R F ≫ whiskerRight ((toPresheaf _).map ψ') _ ) : ψ = ψ' := + freeHomEquiv.injective h + +variable (R) in +/-- The free presheaf of modules functor is left adjoint to the forget functor +`PresheafOfModules.{u} R ⥤ Cᵒᵖ ⥤ Type u`. -/ +noncomputable def freeAdjunction : + free.{u} R ⊣ (toPresheaf R ⋙ (whiskeringRight _ _ _).obj (forget Ab)) := + Adjunction.mkOfHomEquiv + { homEquiv := fun _ _ ↦ freeHomEquiv + homEquiv_naturality_left_symm := fun {F₁ F₂ G} f g ↦ + free_hom_ext (by ext; simp [freeHomEquiv, toPresheaf]) + homEquiv_naturality_right := fun {F G₁ G₂} f g ↦ rfl } + +variable (F G) in +@[simp] +lemma freeAdjunction_homEquiv : (freeAdjunction R).homEquiv F G = freeHomEquiv := by + simp [freeAdjunction, Adjunction.mkOfHomEquiv_homEquiv] + +variable (R F) in +@[simp] +lemma freeAdjunction_unit_app : + (freeAdjunction R).unit.app F = freeAdjunctionUnit R F := rfl + +end + +end PresheafOfModules diff --git a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafification.lean b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafification.lean index db55107701398..55d42463d8a56 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafification.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafification.lean @@ -124,7 +124,7 @@ noncomputable def sheafificationAdjunction : apply (toPresheaf _).map_injective erw [toPresheaf_map_sheafificationHomEquiv] } -lemma sheaififcationAdjunction_homEquiv_apply {P : PresheafOfModules.{v} R₀} +lemma sheafificationAdjunction_homEquiv_apply {P : PresheafOfModules.{v} R₀} {F : SheafOfModules.{v} R} (f : (sheafification α).obj P ⟶ F) : (sheafificationAdjunction α).homEquiv P F f = sheafificationHomEquiv α f := rfl diff --git a/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean b/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean index b43c932f7d3bc..d067b4a921437 100644 --- a/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean +++ b/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean @@ -563,8 +563,7 @@ theorem forget₂ToModuleCatHomotopyEquiv_f_0_eq : · dsimp only [HomotopyEquiv.ofIso, compForgetAugmentedIso] simp only [Iso.symm_hom, eqToIso.inv, HomologicalComplex.eqToHom_f, eqToHom_refl] trans (linearCombination _ fun _ => (1 : k)).comp ((ModuleCat.free k).map (terminal.from _)) - · dsimp - erw [Finsupp.lmapDomain_linearCombination (α := Fin 1 → G) (R := k) (α' := ⊤_ Type u) + · erw [Finsupp.lmapDomain_linearCombination (α := Fin 1 → G) (R := k) (α' := ⊤_ Type u) (v := fun _ => (1 : k)) (v' := fun _ => (1 : k)) (terminal.from ((classifyingSpaceUniversalCover G).obj (Opposite.op (SimplexCategory.mk 0))).V) From e82d87a5d52effc6ba6bd8c1ab2612c5420316b4 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Thu, 24 Oct 2024 09:37:24 +0000 Subject: [PATCH 361/505] feat: LoC mode for `lake exe pole` (#18157) Enables `lake exe pole --loc` which uses lines of code as a quick proxy for the instruction count. This allows running `lake exe pole` on commits that have not been through the speed center, or in cases where it's convenient to avoid the speed center. --- LongestPole/Main.lean | 27 +++++++++++++++++++++++---- 1 file changed, 23 insertions(+), 4 deletions(-) diff --git a/LongestPole/Main.lean b/LongestPole/Main.lean index b4abb23680785..ad483201ffeea 100644 --- a/LongestPole/Main.lean +++ b/LongestPole/Main.lean @@ -54,7 +54,7 @@ def RunResponse.instructions (response : RunResponse) : for m in response.run.result.measurements do let n := m.dimension.benchmark if n.startsWith "~" then - r := r.insert (n.drop 1).toName m.value + r := r.insert (n.drop 1).toName (m.value/10^6) return r def instructions (run : String) : IO (NameMap Float) := @@ -116,6 +116,19 @@ def Float.toStringDecimals (r : Float) (digits : Nat) : String := | [a, b] => a ++ "." ++ b.take digits | _ => r.toString +open System in +-- Lines of code is obviously a `Nat` not a `Float`, +-- but we're using it here as a very rough proxy for instruction count. +def countLOC (modules : List Name) : IO (NameMap Float) := do + let mut r := {} + for m in modules do + let fp := FilePath.mk ((← findOLean m).toString.replace ".lake/build/lib/" "") + |>.withExtension "lean" + if ← fp.pathExists then + let src ← IO.FS.readFile fp + r := r.insert m (src.toList.count '\n').toFloat + return r + /-- Implementation of the longest pole command line program. -/ def longestPoleCLI (args : Cli.Parsed) : IO UInt32 := do let to ← match args.flag? "to" with @@ -126,7 +139,10 @@ def longestPoleCLI (args : Cli.Parsed) : IO UInt32 := do let graph := env.importGraph let sha ← headSha IO.eprintln s!"Analyzing {to} at {sha}" - let instructions ← SpeedCenterAPI.instructions (sha) + let instructions ← if args.hasFlag "loc" then + countLOC (graph.toList.map (·.1)) + else + SpeedCenterAPI.instructions sha let cumulative := cumulativeInstructions instructions graph let total := totalInstructions instructions graph let slowest := slowestParents cumulative graph @@ -138,10 +154,12 @@ def longestPoleCLI (args : Cli.Parsed) : IO UInt32 := do let c := cumulative.find! n' let t := total.find! n' let r := (t / c).toStringDecimals 2 - table := table.push #[n.get!.toString, toString (i/10^6 |>.toUInt64), toString (c/10^6 |>.toUInt64), r] + table := table.push + #[n.get!.toString, toString (i |>.toUInt64), toString (c |>.toUInt64), r] n := slowest.find? n' + let instructionsHeader := if args.hasFlag "loc" then "LoC" else "instructions" IO.println (formatTable - #["file", "instructions", "cumulative", "parallelism"] + #["file", instructionsHeader, "cumulative", "parallelism"] table #[Alignment.left, Alignment.right, Alignment.right, Alignment.center]) return 0 @@ -157,6 +175,7 @@ def pole : Cmd := `[Cli| FLAGS: to : ModuleName; "Calculate the longest pole to the specified module." + loc; "Use lines of code instead of speedcenter instruction counts." ] /-- `lake exe pole` -/ From 30d27d235c72252919d9257ea48721b86cadd777 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Thu, 24 Oct 2024 09:58:17 +0000 Subject: [PATCH 362/505] chore(NumberTheory/Harmonic): shake imports (#18158) This file seems to contain many unused imports. Some of them were around from the start of the file, and some were added in #11750 for reasons I don't entirely understand. (Context: I'm checking which `Defs.lean` files actually only provide definitions.) --- Mathlib/NumberTheory/Harmonic/Defs.lean | 3 --- scripts/noshake.json | 3 --- 2 files changed, 6 deletions(-) diff --git a/Mathlib/NumberTheory/Harmonic/Defs.lean b/Mathlib/NumberTheory/Harmonic/Defs.lean index 1169d86a0e470..effef54ba3920 100644 --- a/Mathlib/NumberTheory/Harmonic/Defs.lean +++ b/Mathlib/NumberTheory/Harmonic/Defs.lean @@ -4,9 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Koundinya Vajjha, Thomas Browning -/ import Mathlib.Algebra.BigOperators.Intervals -import Mathlib.Algebra.Order.BigOperators.Ring.Finset -import Mathlib.Algebra.Order.Field.Basic -import Mathlib.Tactic.Linarith import Mathlib.Tactic.Positivity.Finset /-! diff --git a/scripts/noshake.json b/scripts/noshake.json index 22149664826c0..9cf7d5e4a2312 100644 --- a/scripts/noshake.json +++ b/scripts/noshake.json @@ -326,9 +326,6 @@ ["Mathlib.Control.Traversable.Instances"], "Mathlib.Order.Defs": ["Batteries.Tactic.Trans"], "Mathlib.Order.Basic": ["Batteries.Tactic.Classical"], - "Mathlib.NumberTheory.Harmonic.Defs": - ["Mathlib.Algebra.Order.BigOperators.Ring.Finset", - "Mathlib.Algebra.Order.Field.Basic"], "Mathlib.NumberTheory.Cyclotomic.Basic": ["Mathlib.Init.Core"], "Mathlib.NumberTheory.ArithmeticFunction": ["Mathlib.Tactic.ArithMult"], "Mathlib.MeasureTheory.MeasurableSpace.Defs": ["Mathlib.Tactic.FunProp.Attr"], From a61fcb1bf59f66b04e2d6fa973920190da960526 Mon Sep 17 00:00:00 2001 From: Xavier Roblot Date: Thu, 24 Oct 2024 10:17:22 +0000 Subject: [PATCH 363/505] feat(NumberField/CanonicalEmbedding): prove that `idealLattice` is a `ZLattice` (#18130) This PR is part of the proof of the Analytic Class Number Formula. --- .../NumberField/CanonicalEmbedding/Basic.lean | 48 +++++++++++++++++-- 1 file changed, 43 insertions(+), 5 deletions(-) diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean index 73c8984828756..f119b3f3072eb 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean @@ -137,7 +137,7 @@ theorem latticeBasis_apply [NumberField K] (i : Free.ChooseBasisIndex ℤ (𝓞 simp only [latticeBasis, integralBasis_apply, coe_basisOfLinearIndependentOfCardEqFinrank, Function.comp_apply, Equiv.apply_symm_apply] -theorem mem_span_latticeBasis [NumberField K] (x : (K →+* ℂ) → ℂ) : +theorem mem_span_latticeBasis [NumberField K] {x : (K →+* ℂ) → ℂ} : x ∈ Submodule.span ℤ (Set.range (latticeBasis K)) ↔ x ∈ ((canonicalEmbedding K).comp (algebraMap (𝓞 K) K)).range := by rw [show Set.range (latticeBasis K) = @@ -601,9 +601,9 @@ theorem latticeBasis_apply (i : ChooseBasisIndex ℤ (𝓞 K)) : simp only [latticeBasis, coe_basisOfLinearIndependentOfCardEqFinrank, Function.comp_apply, canonicalEmbedding.latticeBasis_apply, integralBasis_apply, commMap_canonical_eq_mixed] -theorem mem_span_latticeBasis (x : (mixedSpace K)) : +theorem mem_span_latticeBasis {x : (mixedSpace K)} : x ∈ Submodule.span ℤ (Set.range (latticeBasis K)) ↔ - x ∈ ((mixedEmbedding K).comp (algebraMap (𝓞 K) K)).range := by + x ∈ mixedEmbedding.integerLattice K := by rw [show Set.range (latticeBasis K) = (mixedEmbedding K).toIntAlgHom.toLinearMap '' (Set.range (integralBasis K)) by rw [← Set.range_comp]; exact congrArg Set.range (funext (fun i => latticeBasis_apply K i))] @@ -614,7 +614,7 @@ theorem mem_span_latticeBasis (x : (mixedSpace K)) : theorem span_latticeBasis : Submodule.span ℤ (Set.range (latticeBasis K)) = mixedEmbedding.integerLattice K := - Submodule.ext_iff.mpr (mem_span_latticeBasis K) + Submodule.ext_iff.mpr fun _ ↦ mem_span_latticeBasis K instance : DiscreteTopology (mixedEmbedding.integerLattice K) := by classical @@ -626,6 +626,13 @@ instance : IsZLattice ℝ (mixedEmbedding.integerLattice K) := by simp_rw [← span_latticeBasis] exact ZSpan.isZLattice (latticeBasis K) +open Classical in +theorem fundamentalDomain_integerLattice : + MeasureTheory.IsAddFundamentalDomain (mixedEmbedding.integerLattice K) + (ZSpan.fundamentalDomain (latticeBasis K)) := by + rw [← span_latticeBasis] + exact ZSpan.isAddFundamentalDomain (latticeBasis K) _ + theorem mem_rat_span_latticeBasis (x : K) : mixedEmbedding K x ∈ Submodule.span ℚ (Set.range (latticeBasis K)) := by rw [← Basis.sum_repr (integralBasis K) x, map_sum] @@ -652,6 +659,14 @@ theorem latticeBasis_repr_apply (x : K) (i : ChooseBasisIndex ℤ (𝓞 K)) : variable (I : (FractionalIdeal (𝓞 K)⁰ K)ˣ) +/-- The image of the fractional ideal `I` in the mixed space. -/ +abbrev idealLattice : Submodule ℤ (mixedSpace K) := LinearMap.range <| + (mixedEmbedding K).toIntAlgHom.toLinearMap ∘ₗ ((I : Submodule (𝓞 K) K).subtype.restrictScalars ℤ) + +theorem mem_idealLattice {x : mixedSpace K} : + x ∈ idealLattice K I ↔ ∃ y, y ∈ (I : Set K) ∧ mixedEmbedding K y = x := by + simp [idealLattice] + /-- The generalized index of the lattice generated by `I` in the lattice generated by `𝓞 K` is equal to the norm of the ideal `I`. The result is stated in terms of base change determinant and is the translation of `NumberField.det_basisOfFractionalIdeal_eq_absNorm` in @@ -694,7 +709,7 @@ theorem fractionalIdealLatticeBasis_apply (i : ChooseBasisIndex ℤ I) : simp only [fractionalIdealLatticeBasis, Basis.coe_reindex, Basis.coe_mk, Function.comp_apply, Equiv.apply_symm_apply] -theorem mem_span_fractionalIdealLatticeBasis (x : (mixedSpace K)) : +theorem mem_span_fractionalIdealLatticeBasis {x : (mixedSpace K)} : x ∈ Submodule.span ℤ (Set.range (fractionalIdealLatticeBasis K I)) ↔ x ∈ mixedEmbedding K '' I := by rw [show Set.range (fractionalIdealLatticeBasis K I) = @@ -706,6 +721,29 @@ theorem mem_span_fractionalIdealLatticeBasis (x : (mixedSpace K)) : ext; erw [mem_span_basisOfFractionalIdeal]] rfl +theorem span_idealLatticeBasis : + (Submodule.span ℤ (Set.range (fractionalIdealLatticeBasis K I))) = + (mixedEmbedding.idealLattice K I) := by + ext x + simp [mem_span_fractionalIdealLatticeBasis] + +instance : DiscreteTopology (mixedEmbedding.idealLattice K I) := by + classical + rw [← span_idealLatticeBasis] + infer_instance + +open Classical in +instance : IsZLattice ℝ (mixedEmbedding.idealLattice K I) := by + simp_rw [← span_idealLatticeBasis] + exact ZSpan.isZLattice (fractionalIdealLatticeBasis K I) + +open Classical in +theorem fundamentalDomain_idealLattice : + MeasureTheory.IsAddFundamentalDomain (mixedEmbedding.idealLattice K I) + (ZSpan.fundamentalDomain (fractionalIdealLatticeBasis K I)) := by + rw [← span_idealLatticeBasis] + exact ZSpan.isAddFundamentalDomain (fractionalIdealLatticeBasis K I) _ + end integerLattice end NumberField.mixedEmbedding From a86c9d3e03a01e49cdf7aad6190c7c56efa6ad94 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Thu, 24 Oct 2024 10:36:55 +0000 Subject: [PATCH 364/505] feat: the monoidal structure on the category of presheaves of modules (#17300) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This contribution was created as part of the AIM workshop "Formalizing algebraic geometry" in June 2024. Co-authored-by: Dagur Asgeirsson Co-authored-by: Jack McKoen Co-authored-by: Joël Riou <37772949+joelriou@users.noreply.github.com> --- Mathlib.lean | 1 + .../Category/ModuleCat/Monoidal/Basic.lean | 41 +++++- .../Algebra/Category/ModuleCat/Presheaf.lean | 15 +- .../Category/ModuleCat/Presheaf/Monoidal.lean | 129 ++++++++++++++++++ .../Monoidal/Internal/Module.lean | 2 +- 5 files changed, 185 insertions(+), 3 deletions(-) create mode 100644 Mathlib/Algebra/Category/ModuleCat/Presheaf/Monoidal.lean diff --git a/Mathlib.lean b/Mathlib.lean index 0511dec58327e..b852c3e52956e 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -116,6 +116,7 @@ import Mathlib.Algebra.Category.ModuleCat.Presheaf.ChangeOfRings import Mathlib.Algebra.Category.ModuleCat.Presheaf.Colimits import Mathlib.Algebra.Category.ModuleCat.Presheaf.Free import Mathlib.Algebra.Category.ModuleCat.Presheaf.Limits +import Mathlib.Algebra.Category.ModuleCat.Presheaf.Monoidal import Mathlib.Algebra.Category.ModuleCat.Presheaf.Pushforward import Mathlib.Algebra.Category.ModuleCat.Presheaf.Sheafification import Mathlib.Algebra.Category.ModuleCat.Presheaf.Sheafify diff --git a/Mathlib/Algebra/Category/ModuleCat/Monoidal/Basic.lean b/Mathlib/Algebra/Category/ModuleCat/Monoidal/Basic.lean index cad418c48ab2a..b801c34fa8861 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Monoidal/Basic.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Monoidal/Basic.lean @@ -203,10 +203,13 @@ instance : CommRing ((𝟙_ (ModuleCat.{u} R) : ModuleCat.{u} R) : Type u) := namespace MonoidalCategory @[simp] -theorem hom_apply {K L M N : ModuleCat.{u} R} (f : K ⟶ L) (g : M ⟶ N) (k : K) (m : M) : +theorem tensorHom_tmul {K L M N : ModuleCat.{u} R} (f : K ⟶ L) (g : M ⟶ N) (k : K) (m : M) : (f ⊗ g) (k ⊗ₜ m) = f k ⊗ₜ g m := rfl +@[deprecated (since := "2024-09-30")] alias hom_apply := tensorHom_tmul + + @[simp] theorem whiskerLeft_apply (L : ModuleCat.{u} R) {M N : ModuleCat.{u} R} (f : M ⟶ N) (l : L) (m : M) : @@ -249,6 +252,42 @@ theorem associator_inv_apply {M N K : ModuleCat.{u} R} (m : M) (n : N) (k : K) : ((α_ M N K).inv : M ⊗ N ⊗ K ⟶ (M ⊗ N) ⊗ K) (m ⊗ₜ (n ⊗ₜ k)) = m ⊗ₜ n ⊗ₜ k := rfl +variable {M₁ M₂ M₃ M₄ : ModuleCat.{u} R} + +section + +variable (f : M₁ → M₂ → M₃) (h₁ : ∀ m₁ m₂ n, f (m₁ + m₂) n = f m₁ n + f m₂ n) + (h₂ : ∀ (a : R) m n, f (a • m) n = a • f m n) + (h₃ : ∀ m n₁ n₂, f m (n₁ + n₂) = f m n₁ + f m n₂) + (h₄ : ∀ (a : R) m n, f m (a • n) = a • f m n) + +/-- Construct for morphisms from the tensor product of two objects in `ModuleCat`. -/ +noncomputable def tensorLift : M₁ ⊗ M₂ ⟶ M₃ := + TensorProduct.lift (LinearMap.mk₂ R f h₁ h₂ h₃ h₄) + +@[simp] +lemma tensorLift_tmul (m : M₁) (n : M₂) : + tensorLift f h₁ h₂ h₃ h₄ (m ⊗ₜ n) = f m n := rfl + +end + +lemma tensor_ext {f g : M₁ ⊗ M₂ ⟶ M₃} (h : ∀ m n, f (m ⊗ₜ n) = g (m ⊗ₜ n)) : + f = g := + TensorProduct.ext (by ext; apply h) + +/-- Extensionality lemma for morphisms from a module of the form `(M₁ ⊗ M₂) ⊗ M₃`. -/ +lemma tensor_ext₃' {f g : (M₁ ⊗ M₂) ⊗ M₃ ⟶ M₄} + (h : ∀ m₁ m₂ m₃, f (m₁ ⊗ₜ m₂ ⊗ₜ m₃) = g (m₁ ⊗ₜ m₂ ⊗ₜ m₃)) : + f = g := + TensorProduct.ext_threefold h + +/-- Extensionality lemma for morphisms from a module of the form `M₁ ⊗ (M₂ ⊗ M₃)`. -/ +lemma tensor_ext₃ {f g : M₁ ⊗ (M₂ ⊗ M₃) ⟶ M₄} + (h : ∀ m₁ m₂ m₃, f (m₁ ⊗ₜ (m₂ ⊗ₜ m₃)) = g (m₁ ⊗ₜ (m₂ ⊗ₜ m₃))) : + f = g := by + rw [← cancel_epi (α_ _ _ _).hom] + exact tensor_ext₃' h + end MonoidalCategory open Opposite diff --git a/Mathlib/Algebra/Category/ModuleCat/Presheaf.lean b/Mathlib/Algebra/Category/ModuleCat/Presheaf.lean index caf2f760e93ea..653280c6b57bd 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Presheaf.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Presheaf.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2023 Kim Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Kim Morrison +Authors: Kim Morrison, Joël Riou -/ import Mathlib.Algebra.Category.ModuleCat.ChangeOfRings import Mathlib.Algebra.Category.Ring.Basic @@ -96,6 +96,19 @@ lemma naturality_apply (f : M₁ ⟶ M₂) {X Y : Cᵒᵖ} (g : X ⟶ Y) (x : M Hom.app f Y (M₁.map g x) = M₂.map g (Hom.app f X x) := congr_fun ((forget _).congr_map (Hom.naturality f g)) x +/-- Constructor for isomorphisms in the category of presheaves of modules. -/ +@[simps!] +def isoMk (app : ∀ (X : Cᵒᵖ), M₁.obj X ≅ M₂.obj X) + (naturality : ∀ ⦃X Y : Cᵒᵖ⦄ (f : X ⟶ Y), + M₁.map f ≫ (ModuleCat.restrictScalars (R.map f)).map (app Y).hom = + (app X).hom ≫ M₂.map f := by aesop_cat) : M₁ ≅ M₂ where + hom := { app := fun X ↦ (app X).hom } + inv := + { app := fun X ↦ (app X).inv + naturality := fun {X Y} f ↦ by + rw [← cancel_epi (app X).hom, ← reassoc_of% (naturality f), Iso.map_hom_inv_id, + Category.comp_id, Iso.hom_inv_id_assoc]} + /-- The underlying presheaf of abelian groups of a presheaf of modules. -/ def presheaf : Cᵒᵖ ⥤ Ab where obj X := (forget₂ _ _).obj (M.obj X) diff --git a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Monoidal.lean b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Monoidal.lean new file mode 100644 index 0000000000000..2b9de6dc8f4e9 --- /dev/null +++ b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Monoidal.lean @@ -0,0 +1,129 @@ +/- +Copyright (c) 2024 Dagur Asgeirsson. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Dagur Asgeirsson, Jack McKoen, Joël Riou +-/ +import Mathlib.Algebra.Category.ModuleCat.Presheaf +import Mathlib.Algebra.Category.ModuleCat.Monoidal.Basic + +/-! +# The monoidal category structure on presheaves of modules + +Given a presheaf of commutative rings `R : Cᵒᵖ ⥤ CommRingCat`, we construct +the monoidal category structure on the category of presheaves of modules +`PresheafOfModules (R ⋙ forget₂ _ _)`. The tensor product `M₁ ⊗ M₂` is defined +as the presheaf of modules which sends `X : Cᵒᵖ` to `M₁.obj X ⊗ M₂.obj X`. + +## Notes + +This contribution was created as part of the AIM workshop +"Formalizing algebraic geometry" in June 2024. + +-/ + +open CategoryTheory MonoidalCategory Category + +universe v u v₁ u₁ + +variable {C : Type*} [Category C] {R : Cᵒᵖ ⥤ CommRingCat.{u}} + +instance (X : Cᵒᵖ) : CommRing ((R ⋙ forget₂ _ RingCat).obj X) := + inferInstanceAs (CommRing (R.obj X)) + +namespace PresheafOfModules + +namespace Monoidal + +variable (M₁ M₂ M₃ M₄ : PresheafOfModules.{u} (R ⋙ forget₂ _ _)) + +/-- Auxiliary definition for `tensorObj`. -/ +noncomputable def tensorObjMap {X Y : Cᵒᵖ} (f : X ⟶ Y) : M₁.obj X ⊗ M₂.obj X ⟶ + (ModuleCat.restrictScalars (R.map f)).obj (M₁.obj Y ⊗ M₂.obj Y) := + ModuleCat.MonoidalCategory.tensorLift (fun m₁ m₂ ↦ M₁.map f m₁ ⊗ₜ M₂.map f m₂) + (by intro m₁ m₁' m₂; dsimp; rw [map_add, TensorProduct.add_tmul]) + (by intro a m₁ m₂; dsimp; erw [M₁.map_smul]; rfl) + (by intro m₁ m₂ m₂'; dsimp; rw [map_add, TensorProduct.tmul_add]) + (by intro a m₁ m₂; dsimp; erw [M₂.map_smul, TensorProduct.tmul_smul (r := R.map f a)]; rfl) + +/-- The tensor product of two presheaves of modules. -/ +@[simps obj] +noncomputable def tensorObj : PresheafOfModules (R ⋙ forget₂ _ _) where + obj X := M₁.obj X ⊗ M₂.obj X + map f := tensorObjMap M₁ M₂ f + map_id X := ModuleCat.MonoidalCategory.tensor_ext (by + intro m₁ m₂ + dsimp [tensorObjMap] + simp only [map_id, Functor.comp_obj, CommRingCat.forgetToRingCat_obj, Functor.comp_map, + ModuleCat.restrictScalarsId'_inv_app, ModuleCat.restrictScalarsId'App_inv_apply] + rfl) + map_comp f g := ModuleCat.MonoidalCategory.tensor_ext (by + intro m₁ m₂ + dsimp [tensorObjMap] + simp only [map_comp, Functor.comp_obj, CommRingCat.forgetToRingCat_obj, Functor.comp_map, + ModuleCat.restrictScalarsComp'_inv_app, ModuleCat.coe_comp, Function.comp_apply, + ModuleCat.restrictScalars.map_apply, ModuleCat.restrictScalarsComp'App_inv_apply] + rfl) + +variable {M₁ M₂ M₃ M₄} + +@[simp] +lemma tensorObj_map_tmul {X Y : Cᵒᵖ} (f : X ⟶ Y) (m₁ : M₁.obj X) (m₂ : M₂.obj X) : + DFunLike.coe (α := (M₁.obj X ⊗ M₂.obj X : _)) + (β := fun _ ↦ (ModuleCat.restrictScalars + ((forget₂ CommRingCat RingCat).map (R.map f))).obj (M₁.obj Y ⊗ M₂.obj Y)) + ((tensorObj M₁ M₂).map f) (m₁ ⊗ₜ[R.obj X] m₂) = M₁.map f m₁ ⊗ₜ[R.obj Y] M₂.map f m₂ := rfl + +/-- The tensor product of two morphisms of presheaves of modules. -/ +@[simps] +noncomputable def tensorHom (f : M₁ ⟶ M₂) (g : M₃ ⟶ M₄) : tensorObj M₁ M₃ ⟶ tensorObj M₂ M₄ where + app X := f.app X ⊗ g.app X + naturality {X Y} φ := ModuleCat.MonoidalCategory.tensor_ext (fun m₁ m₃ ↦ by + dsimp + rw [tensorObj_map_tmul] + erw [ModuleCat.MonoidalCategory.tensorHom_tmul, tensorObj_map_tmul, + naturality_apply, naturality_apply] + rfl) + +end Monoidal + +open Monoidal + +open ModuleCat.MonoidalCategory in +noncomputable instance monoidalCategoryStruct : + MonoidalCategoryStruct (PresheafOfModules.{u} (R ⋙ forget₂ _ _)) where + tensorObj := tensorObj + whiskerLeft _ _ _ g := tensorHom (𝟙 _) g + whiskerRight f _ := tensorHom f (𝟙 _) + tensorHom := tensorHom + tensorUnit := unit _ + associator M₁ M₂ M₃ := isoMk (fun _ ↦ α_ _ _ _) + (fun _ _ _ ↦ ModuleCat.MonoidalCategory.tensor_ext₃' (by intros; rfl)) + leftUnitor M := Iso.symm (isoMk (fun _ ↦ (λ_ _).symm) (fun X Y f ↦ by + ext m + dsimp + erw [leftUnitor_inv_apply, leftUnitor_inv_apply, tensorObj_map_tmul, (R.map f).map_one] + rfl)) + rightUnitor M := Iso.symm (isoMk (fun _ ↦ (ρ_ _).symm) (fun X Y f ↦ by + ext m + dsimp + erw [rightUnitor_inv_apply, rightUnitor_inv_apply, tensorObj_map_tmul, (R.map f).map_one] + rfl)) + +noncomputable instance monoidalCategory : + MonoidalCategory (PresheafOfModules.{u} (R ⋙ forget₂ _ _)) where + tensorHom_def _ _ := by ext1; apply tensorHom_def + tensor_id _ _ := by ext1; apply tensor_id + tensor_comp _ _ _ _ := by ext1; apply tensor_comp + whiskerLeft_id M₁ M₂ := by + ext1 X + apply MonoidalCategory.whiskerLeft_id (C := ModuleCat (R.obj X)) + id_whiskerRight _ _ := by + ext1 X + apply MonoidalCategory.id_whiskerRight (C := ModuleCat (R.obj X)) + associator_naturality _ _ _ := by ext1; apply associator_naturality + leftUnitor_naturality _ := by ext1; apply leftUnitor_naturality + rightUnitor_naturality _ := by ext1; apply rightUnitor_naturality + pentagon _ _ _ _ := by ext1; apply pentagon + triangle _ _ := by ext1; apply triangle + +end PresheafOfModules diff --git a/Mathlib/CategoryTheory/Monoidal/Internal/Module.lean b/Mathlib/CategoryTheory/Monoidal/Internal/Module.lean index b579897c35ad0..65be75a189756 100644 --- a/Mathlib/CategoryTheory/Monoidal/Internal/Module.lean +++ b/Mathlib/CategoryTheory/Monoidal/Internal/Module.lean @@ -135,7 +135,7 @@ def inverseObj (A : AlgebraCat.{u} R) : Mon_ (ModuleCat.{u} R) where refine TensorProduct.ext <| TensorProduct.ext <| LinearMap.ext fun x => LinearMap.ext fun y => LinearMap.ext fun z => ?_ dsimp only [AlgebraCat.id_apply, TensorProduct.mk_apply, LinearMap.compr₂_apply, - Function.comp_apply, ModuleCat.MonoidalCategory.hom_apply, AlgebraCat.coe_comp, + Function.comp_apply, ModuleCat.MonoidalCategory.tensorHom_tmul, AlgebraCat.coe_comp, MonoidalCategory.associator_hom_apply] rw [compr₂_apply, compr₂_apply] -- This used to be `rw`, but we need `erw` after leanprover/lean4#2644 From a5a07115f62b708edf3fbf128cc60041492cf418 Mon Sep 17 00:00:00 2001 From: Xavier Roblot Date: Thu, 24 Oct 2024 10:44:03 +0000 Subject: [PATCH 365/505] feat(RingTheory/Ideal/Norm): Generalize `Ideal.finite_setOf_absNorm_eq` and add `Ideal.finite_setOf_absNorm_le` (#18132) Also: - prove some results about the norm of `nonZeroDivisors` ideals - golf the proof of Dirichlet unit theorem using the generalization - clean up the import of `RingTheory.Ideal.Norm` This PR is part of the proof of the Analytic Class Number Formula. --- .../NumberField/Units/DirichletTheorem.lean | 27 +++---- Mathlib/RingTheory/Ideal/Norm.lean | 78 ++++++++++++++----- 2 files changed, 67 insertions(+), 38 deletions(-) diff --git a/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean b/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean index cc34b948b40e0..068bff159224a 100644 --- a/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean +++ b/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean @@ -250,10 +250,6 @@ def seq : ℕ → { x : 𝓞 K // x ≠ 0 } theorem seq_ne_zero (n : ℕ) : algebraMap (𝓞 K) K (seq K w₁ hB n) ≠ 0 := RingOfIntegers.coe_ne_zero_iff.mpr (seq K w₁ hB n).prop -/-- The terms of the sequence have nonzero norm. -/ -theorem seq_norm_ne_zero (n : ℕ) : Algebra.norm ℤ (seq K w₁ hB n : 𝓞 K) ≠ 0 := - Algebra.norm_ne_zero_iff.mpr (Subtype.coe_ne_coe.1 (seq_ne_zero K w₁ hB n)) - /-- The sequence is strictly decreasing at infinite places distinct from `w₁`. -/ theorem seq_decreasing {n m : ℕ} (h : n < m) (w : InfinitePlace K) (hw : w ≠ w₁) : w (algebraMap (𝓞 K) K (seq K w₁ hB m)) < w (algebraMap (𝓞 K) K (seq K w₁ hB n)) := by @@ -301,20 +297,15 @@ theorem exists_unit (w₁ : InfinitePlace K) : _ = w (algebraMap (𝓞 K) K (seq K w₁ hB m) * (algebraMap (𝓞 K) K (seq K w₁ hB n))⁻¹) := by rw [← congr_arg (algebraMap (𝓞 K) K) hu.choose_spec, mul_comm, map_mul (algebraMap _ _), ← mul_assoc, inv_mul_cancel₀ (seq_ne_zero K w₁ hB n), one_mul] - _ = w (algebraMap (𝓞 K) K (seq K w₁ hB m)) * w (algebraMap (𝓞 K) K (seq K w₁ hB n))⁻¹ := - _root_.map_mul _ _ _ - _ < 1 := by - rw [map_inv₀, mul_inv_lt_iff₀ (pos_iff.mpr (seq_ne_zero K w₁ hB n)), one_mul] - exact seq_decreasing K w₁ hB hnm w hw - refine Set.Finite.exists_lt_map_eq_of_forall_mem - (t := { I : Ideal (𝓞 K) | 1 ≤ Ideal.absNorm I ∧ Ideal.absNorm I ≤ B }) - (fun n => ?_) ?_ - · rw [Set.mem_setOf_eq, Ideal.absNorm_span_singleton] - refine ⟨?_, seq_norm_le K w₁ hB n⟩ - exact Nat.one_le_iff_ne_zero.mpr (Int.natAbs_ne_zero.mpr (seq_norm_ne_zero K w₁ hB n)) - · rw [show { I : Ideal (𝓞 K) | 1 ≤ Ideal.absNorm I ∧ Ideal.absNorm I ≤ B } = - (⋃ n ∈ Set.Icc 1 B, { I : Ideal (𝓞 K) | Ideal.absNorm I = n }) by ext; simp] - exact Set.Finite.biUnion (Set.finite_Icc _ _) (fun n hn => Ideal.finite_setOf_absNorm_eq hn.1) + _ = w (algebraMap (𝓞 K) K (seq K w₁ hB m)) * w (algebraMap (𝓞 K) K (seq K w₁ hB n))⁻¹ := + _root_.map_mul _ _ _ + _ < 1 := by + rw [map_inv₀, mul_inv_lt_iff₀' (pos_iff.mpr (seq_ne_zero K w₁ hB n)), mul_one] + exact seq_decreasing K w₁ hB hnm w hw + refine Set.Finite.exists_lt_map_eq_of_forall_mem (t := {I : Ideal (𝓞 K) | Ideal.absNorm I ≤ B}) + (fun n ↦ ?_) (Ideal.finite_setOf_absNorm_le B) + rw [Set.mem_setOf_eq, Ideal.absNorm_span_singleton] + exact seq_norm_le K w₁ hB n theorem unitLattice_span_eq_top : Submodule.span ℝ (unitLattice K : Set ({w : InfinitePlace K // w ≠ w₀} → ℝ)) = ⊤ := by diff --git a/Mathlib/RingTheory/Ideal/Norm.lean b/Mathlib/RingTheory/Ideal/Norm.lean index 6e1a0aa1de64f..ae759f8605dbe 100644 --- a/Mathlib/RingTheory/Ideal/Norm.lean +++ b/Mathlib/RingTheory/Ideal/Norm.lean @@ -4,15 +4,11 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Anne Baanen, Alex J. Best -/ import Mathlib.Algebra.CharP.Quotient -import Mathlib.Algebra.GroupWithZero.NonZeroDivisors -import Mathlib.Data.Finsupp.Fintype import Mathlib.Data.Int.AbsoluteValue import Mathlib.Data.Int.Associated import Mathlib.LinearAlgebra.FreeModule.Determinant import Mathlib.LinearAlgebra.FreeModule.IdealQuotient import Mathlib.RingTheory.DedekindDomain.PID -import Mathlib.RingTheory.Ideal.Basis -import Mathlib.RingTheory.LocalProperties.Basic import Mathlib.RingTheory.Localization.NormTrace /-! @@ -391,24 +387,66 @@ theorem absNorm_eq_zero_iff {I : Ideal S} : Ideal.absNorm I = 0 ↔ I = ⊥ := b · rintro rfl exact absNorm_bot -theorem absNorm_ne_zero_of_nonZeroDivisors (I : (Ideal S)⁰) : Ideal.absNorm (I : Ideal S) ≠ 0 := - Ideal.absNorm_eq_zero_iff.not.mpr <| nonZeroDivisors.coe_ne_zero _ +theorem absNorm_ne_zero_iff_mem_nonZeroDivisors {I : Ideal S} : + absNorm I ≠ 0 ↔ I ∈ (Ideal S)⁰ := by + simp_rw [ne_eq, Ideal.absNorm_eq_zero_iff, mem_nonZeroDivisors_iff_ne_zero, Submodule.zero_eq_bot] -theorem finite_setOf_absNorm_eq [CharZero S] {n : ℕ} (hn : 0 < n) : +theorem absNorm_pos_iff_mem_nonZeroDivisors {I : Ideal S} : + 0 < absNorm I ↔ I ∈ (Ideal S)⁰ := by + rw [← absNorm_ne_zero_iff_mem_nonZeroDivisors, Nat.pos_iff_ne_zero] + +theorem absNorm_ne_zero_of_nonZeroDivisors (I : (Ideal S)⁰) : absNorm (I : Ideal S) ≠ 0 := + absNorm_ne_zero_iff_mem_nonZeroDivisors.mpr (SetLike.coe_mem I) + +theorem absNorm_pos_of_nonZeroDivisors (I : (Ideal S)⁰) : 0 < absNorm (I : Ideal S) := + absNorm_pos_iff_mem_nonZeroDivisors.mpr (SetLike.coe_mem I) + +theorem finite_setOf_absNorm_eq [CharZero S] (n : ℕ) : {I : Ideal S | Ideal.absNorm I = n}.Finite := by - let f := fun I : Ideal S => Ideal.map (Ideal.Quotient.mk (@Ideal.span S _ {↑n})) I - refine @Set.Finite.of_finite_image _ _ _ f ?_ ?_ - · suffices Finite (S ⧸ @Ideal.span S _ {↑n}) by - let g := ((↑) : Ideal (S ⧸ @Ideal.span S _ {↑n}) → Set (S ⧸ @Ideal.span S _ {↑n})) - refine @Set.Finite.of_finite_image _ _ _ g ?_ SetLike.coe_injective.injOn - exact Set.Finite.subset (@Set.finite_univ _ (@Set.finite' _ this)) (Set.subset_univ _) - rw [← absNorm_ne_zero_iff, absNorm_span_singleton] - simpa only [Ne, Int.natAbs_eq_zero, Algebra.norm_eq_zero_iff, Nat.cast_eq_zero] using - ne_of_gt hn - · intro I hI J hJ h - rw [← comap_map_mk (span_singleton_absNorm_le I), ← hI.symm, ← - comap_map_mk (span_singleton_absNorm_le J), ← hJ.symm] - congr + obtain hn | hn := Nat.eq_zero_or_pos n + · simp only [hn, absNorm_eq_zero_iff, Set.setOf_eq_eq_singleton, Set.finite_singleton] + · let f := fun I : Ideal S => Ideal.map (Ideal.Quotient.mk (@Ideal.span S _ {↑n})) I + refine Set.Finite.of_finite_image (f := f) ?_ ?_ + · suffices Finite (S ⧸ @Ideal.span S _ {↑n}) by + let g := ((↑) : Ideal (S ⧸ @Ideal.span S _ {↑n}) → Set (S ⧸ @Ideal.span S _ {↑n})) + refine Set.Finite.of_finite_image (f := g) ?_ SetLike.coe_injective.injOn + exact Set.Finite.subset Set.finite_univ (Set.subset_univ _) + rw [← absNorm_ne_zero_iff, absNorm_span_singleton] + simpa only [Ne, Int.natAbs_eq_zero, Algebra.norm_eq_zero_iff, Nat.cast_eq_zero] using + ne_of_gt hn + · intro I hI J hJ h + rw [← comap_map_mk (span_singleton_absNorm_le I), ← hI.symm, ← + comap_map_mk (span_singleton_absNorm_le J), ← hJ.symm] + congr + +theorem finite_setOf_absNorm_le [CharZero S] (n : ℕ) : + {I : Ideal S | Ideal.absNorm I ≤ n}.Finite := by + rw [show {I : Ideal S | Ideal.absNorm I ≤ n} = + (⋃ i ∈ Set.Icc 0 n, {I : Ideal S | Ideal.absNorm I = i}) by ext; simp] + refine Set.Finite.biUnion (Set.finite_Icc 0 n) (fun i _ => Ideal.finite_setOf_absNorm_eq i) + +theorem card_norm_le_eq_card_norm_le_add_one (n : ℕ) [CharZero S] : + Nat.card {I : Ideal S // absNorm I ≤ n} = + Nat.card {I : (Ideal S)⁰ // absNorm (I : Ideal S) ≤ n} + 1 := by + classical + have : Finite {I : Ideal S // I ∈ (Ideal S)⁰ ∧ absNorm I ≤ n} := + (finite_setOf_absNorm_le n).subset fun _ ⟨_, h⟩ ↦ h + have : Finite {I : Ideal S // I ∉ (Ideal S)⁰ ∧ absNorm I ≤ n} := + (finite_setOf_absNorm_le n).subset fun _ ⟨_, h⟩ ↦ h + rw [Nat.card_congr (Equiv.subtypeSubtypeEquivSubtypeInter (fun I ↦ I ∈ (Ideal S)⁰) + (fun I ↦ absNorm I ≤ n))] + let e : {I : Ideal S // absNorm I ≤ n} ≃ {I : Ideal S // I ∈ (Ideal S)⁰ ∧ absNorm I ≤ n} ⊕ + {I : Ideal S // I ∉ (Ideal S)⁰ ∧ absNorm I ≤ n} := by + refine (Equiv.subtypeEquivRight ?_).trans (subtypeOrEquiv _ _ ?_) + · intro _ + simp_rw [← or_and_right, em, true_and] + · exact Pi.disjoint_iff.mpr fun I ↦ Prop.disjoint_iff.mpr (by tauto) + simp_rw [Nat.card_congr e, Nat.card_sum, add_right_inj] + conv_lhs => + enter [1, 1, I] + rw [← absNorm_ne_zero_iff_mem_nonZeroDivisors, ne_eq, not_not, and_iff_left_iff_imp.mpr + (fun h ↦ by rw [h]; exact Nat.zero_le n), absNorm_eq_zero_iff] + rw [Nat.card_unique] theorem norm_dvd_iff {x : S} (hx : Prime (Algebra.norm ℤ x)) {y : ℤ} : Algebra.norm ℤ x ∣ y ↔ x ∣ y := by From 0a0178f5f7dfe519f09d3d06cf6b1179f404f2b2 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Thu, 24 Oct 2024 10:44:05 +0000 Subject: [PATCH 366/505] feat(Algebra/Category/ModuleCat/Presheaf): epi and mono in the category of presheaves of modules (#18159) In the category of presheaves of modules, epi and mono can be characterized in terms of surjective/injective maps. --- Mathlib.lean | 1 + .../Category/ModuleCat/Presheaf/EpiMono.lean | 64 +++++++++++++++++++ 2 files changed, 65 insertions(+) create mode 100644 Mathlib/Algebra/Category/ModuleCat/Presheaf/EpiMono.lean diff --git a/Mathlib.lean b/Mathlib.lean index b852c3e52956e..d1b689b301915 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -114,6 +114,7 @@ import Mathlib.Algebra.Category.ModuleCat.Presheaf import Mathlib.Algebra.Category.ModuleCat.Presheaf.Abelian import Mathlib.Algebra.Category.ModuleCat.Presheaf.ChangeOfRings import Mathlib.Algebra.Category.ModuleCat.Presheaf.Colimits +import Mathlib.Algebra.Category.ModuleCat.Presheaf.EpiMono import Mathlib.Algebra.Category.ModuleCat.Presheaf.Free import Mathlib.Algebra.Category.ModuleCat.Presheaf.Limits import Mathlib.Algebra.Category.ModuleCat.Presheaf.Monoidal diff --git a/Mathlib/Algebra/Category/ModuleCat/Presheaf/EpiMono.lean b/Mathlib/Algebra/Category/ModuleCat/Presheaf/EpiMono.lean new file mode 100644 index 0000000000000..2bdcf0998ad7c --- /dev/null +++ b/Mathlib/Algebra/Category/ModuleCat/Presheaf/EpiMono.lean @@ -0,0 +1,64 @@ +/- +Copyright (c) 2024 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ + +import Mathlib.Algebra.Category.ModuleCat.Presheaf.Colimits +import Mathlib.Algebra.Category.ModuleCat.Presheaf.Limits + +/-! +# Epimorphisms and monomorphisms in the category of presheaves of modules + +In this file, we give characterizations of epimorphisms and monomorphisms +in the category of presheaves of modules. + +-/ + +universe v v₁ u₁ u + +open CategoryTheory + +namespace PresheafOfModules + +variable {C : Type u₁} [Category.{v₁} C] {R : Cᵒᵖ ⥤ RingCat.{u}} + {M₁ M₂ : PresheafOfModules.{v} R} {f : M₁ ⟶ M₂} + +lemma epi_of_surjective (hf : ∀ ⦃X : Cᵒᵖ⦄, Function.Surjective (f.app X)) : Epi f where + left_cancellation g₁ g₂ hg := by + ext X m₂ + obtain ⟨m₁, rfl⟩ := hf m₂ + exact congr_fun ((evaluation R X ⋙ forget _).congr_map hg) m₁ + +lemma mono_of_injective (hf : ∀ ⦃X : Cᵒᵖ⦄, Function.Injective (f.app X)) : Mono f where + right_cancellation {M} g₁ g₂ hg := by + ext X m + exact hf (congr_fun ((evaluation R X ⋙ forget _).congr_map hg) m) + +variable (f) + +instance [Epi f] (X : Cᵒᵖ) : Epi (f.app X) := + inferInstanceAs (Epi ((evaluation R X).map f)) + +instance [Mono f] (X : Cᵒᵖ) : Mono (f.app X) := + inferInstanceAs (Mono ((evaluation R X).map f)) + +lemma surjective_of_epi [Epi f] (X : Cᵒᵖ) : + Function.Surjective (f.app X) := by + rw [← ModuleCat.epi_iff_surjective] + infer_instance + +lemma injective_of_mono [Mono f] (X : Cᵒᵖ) : + Function.Injective (f.app X) := by + rw [← ModuleCat.mono_iff_injective] + infer_instance + +lemma epi_iff_surjective : + Epi f ↔ ∀ ⦃X : Cᵒᵖ⦄, Function.Surjective (f.app X) := + ⟨fun _ ↦ surjective_of_epi f, epi_of_surjective⟩ + +lemma mono_iff_surjective : + Mono f ↔ ∀ ⦃X : Cᵒᵖ⦄, Function.Injective (f.app X) := + ⟨fun _ ↦ injective_of_mono f, mono_of_injective⟩ + +end PresheafOfModules From f7b57bce1a08771bc6b7d35d4275bb6aba0697b5 Mon Sep 17 00:00:00 2001 From: Hannah Scholz Date: Thu, 24 Oct 2024 11:27:38 +0000 Subject: [PATCH 367/505] feat: `closedBall` and `ball` are singletons in `Subsingleton` (#18114) --- Mathlib/Topology/MetricSpace/Pseudo/Defs.lean | 9 +++++++++ 1 file changed, 9 insertions(+) diff --git a/Mathlib/Topology/MetricSpace/Pseudo/Defs.lean b/Mathlib/Topology/MetricSpace/Pseudo/Defs.lean index de82c05c1a324..19689b2f5cdf6 100644 --- a/Mathlib/Topology/MetricSpace/Pseudo/Defs.lean +++ b/Mathlib/Topology/MetricSpace/Pseudo/Defs.lean @@ -401,6 +401,15 @@ theorem sphere_eq_empty_of_subsingleton [Subsingleton α] (hε : ε ≠ 0) : sph instance sphere_isEmpty_of_subsingleton [Subsingleton α] [NeZero ε] : IsEmpty (sphere x ε) := by rw [sphere_eq_empty_of_subsingleton (NeZero.ne ε)]; infer_instance +theorem closedBall_eq_singleton_of_subsingleton [Subsingleton α] (h : 0 ≤ ε) : + closedBall x ε = {x} := by + ext x' + simpa [Subsingleton.allEq x x'] + +theorem ball_eq_singleton_of_subsingleton [Subsingleton α] (h : 0 < ε) : ball x ε = {x} := by + ext x' + simpa [Subsingleton.allEq x x'] + theorem mem_closedBall_self (h : 0 ≤ ε) : x ∈ closedBall x ε := by rwa [mem_closedBall, dist_self] From 29a9413719a817d9137a65157bb7721fe0bd9cf3 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Thu, 24 Oct 2024 11:47:41 +0000 Subject: [PATCH 368/505] feat: `swap` and `compl` of order relations (#18045) --- Mathlib/Logic/Function/Basic.lean | 5 +++++ Mathlib/Order/Basic.lean | 5 +++++ 2 files changed, 10 insertions(+) diff --git a/Mathlib/Logic/Function/Basic.lean b/Mathlib/Logic/Function/Basic.lean index a7cb654f6a4e7..58b832fa43354 100644 --- a/Mathlib/Logic/Function/Basic.lean +++ b/Mathlib/Logic/Function/Basic.lean @@ -65,6 +65,11 @@ lemma funext_iff_of_subsingleton [Subsingleton α] {g : α → β} (x y : α) : · rwa [Subsingleton.elim x z, Subsingleton.elim y z] at h · rw [h, Subsingleton.elim x y] +theorem swap_lt {α} [Preorder α] : swap (· < · : α → α → _) = (· > ·) := rfl +theorem swap_le {α} [Preorder α] : swap (· ≤ · : α → α → _) = (· ≥ ·) := rfl +theorem swap_gt {α} [Preorder α] : swap (· > · : α → α → _) = (· < ·) := rfl +theorem swap_ge {α} [Preorder α] : swap (· ≥ · : α → α → _) = (· ≤ ·) := rfl + protected theorem Bijective.injective {f : α → β} (hf : Bijective f) : Injective f := hf.1 protected theorem Bijective.surjective {f : α → β} (hf : Bijective f) : Surjective f := hf.2 diff --git a/Mathlib/Order/Basic.lean b/Mathlib/Order/Basic.lean index 1249ba8d9278a..fb57460b6d250 100644 --- a/Mathlib/Order/Basic.lean +++ b/Mathlib/Order/Basic.lean @@ -754,6 +754,11 @@ instance IsIrrefl.compl (r) [IsIrrefl α r] : IsRefl α rᶜ := instance IsRefl.compl (r) [IsRefl α r] : IsIrrefl α rᶜ := ⟨fun a ↦ not_not_intro (refl a)⟩ +theorem compl_lt [LinearOrder α] : (· < · : α → α → _)ᶜ = (· ≥ ·) := by ext; simp [compl] +theorem compl_le [LinearOrder α] : (· ≤ · : α → α → _)ᶜ = (· > ·) := by ext; simp [compl] +theorem compl_gt [LinearOrder α] : (· > · : α → α → _)ᶜ = (· ≤ ·) := by ext; simp [compl] +theorem compl_ge [LinearOrder α] : (· ≥ · : α → α → _)ᶜ = (· < ·) := by ext; simp [compl] + /-! ### Order instances on the function space -/ From ba2cab6d66d37935d50184b5732c3e3e01874dc1 Mon Sep 17 00:00:00 2001 From: Markus Himmel Date: Thu, 24 Oct 2024 11:47:42 +0000 Subject: [PATCH 369/505] chore: move CategoryTheory.Comma.Presheaf to CategoryTheory.Comma.Presheaf.Basic (#18162) In preparation for #17451. Co-authored-by: Markus Himmel --- Mathlib.lean | 2 +- .../CategoryTheory/Comma/{Presheaf.lean => Presheaf/Basic.lean} | 0 Mathlib/CategoryTheory/Limits/Presheaf.lean | 2 +- 3 files changed, 2 insertions(+), 2 deletions(-) rename Mathlib/CategoryTheory/Comma/{Presheaf.lean => Presheaf/Basic.lean} (100%) diff --git a/Mathlib.lean b/Mathlib.lean index d1b689b301915..d3acec601583c 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1523,7 +1523,7 @@ import Mathlib.CategoryTheory.CommSq import Mathlib.CategoryTheory.Comma.Arrow import Mathlib.CategoryTheory.Comma.Basic import Mathlib.CategoryTheory.Comma.Over -import Mathlib.CategoryTheory.Comma.Presheaf +import Mathlib.CategoryTheory.Comma.Presheaf.Basic import Mathlib.CategoryTheory.Comma.StructuredArrow import Mathlib.CategoryTheory.ComposableArrows import Mathlib.CategoryTheory.ConcreteCategory.Basic diff --git a/Mathlib/CategoryTheory/Comma/Presheaf.lean b/Mathlib/CategoryTheory/Comma/Presheaf/Basic.lean similarity index 100% rename from Mathlib/CategoryTheory/Comma/Presheaf.lean rename to Mathlib/CategoryTheory/Comma/Presheaf/Basic.lean diff --git a/Mathlib/CategoryTheory/Limits/Presheaf.lean b/Mathlib/CategoryTheory/Limits/Presheaf.lean index 79b189b11e50c..ca7b9175b9037 100644 --- a/Mathlib/CategoryTheory/Limits/Presheaf.lean +++ b/Mathlib/CategoryTheory/Limits/Presheaf.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, Joël Riou -/ -import Mathlib.CategoryTheory.Comma.Presheaf +import Mathlib.CategoryTheory.Comma.Presheaf.Basic import Mathlib.CategoryTheory.Elements import Mathlib.CategoryTheory.Functor.KanExtension.Adjunction import Mathlib.CategoryTheory.Limits.Final From dfb1dd3875cc7faaaed3e4c289836eb138e269b2 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Thu, 24 Oct 2024 12:28:53 +0000 Subject: [PATCH 370/505] chore(CategoryTheory): generalize universes for multiequalizers (#17997) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Multi(co)equalizers involve two index sets. Until now, they were assumed to be in the same universe. In this PR, we allow two distinct universes. This shall ease the formalisation of (co)ends, which require heterogeneous universes. Co-authored-by: Joël Riou <37772949+joelriou@users.noreply.github.com> --- .../Limits/Shapes/ConcreteCategory.lean | 16 ++--- .../Limits/Shapes/Multiequalizer.lean | 60 ++++++++++--------- 2 files changed, 39 insertions(+), 37 deletions(-) diff --git a/Mathlib/CategoryTheory/Limits/Shapes/ConcreteCategory.lean b/Mathlib/CategoryTheory/Limits/Shapes/ConcreteCategory.lean index 6c954e928f31d..d16936f661fa4 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/ConcreteCategory.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/ConcreteCategory.lean @@ -31,7 +31,7 @@ wide-pullbacks, wide-pushouts, multiequalizers and cokernels. -/ -universe w v u t r +universe w w' v u t r namespace CategoryTheory.Limits.Concrete @@ -247,9 +247,9 @@ end WidePullback section Multiequalizer -variable [ConcreteCategory.{max w v} C] +variable [ConcreteCategory.{max w w' v} C] -theorem multiequalizer_ext {I : MulticospanIndex.{w} C} [HasMultiequalizer I] +theorem multiequalizer_ext {I : MulticospanIndex.{w, w'} C} [HasMultiequalizer I] [PreservesLimit I.multicospan (forget C)] (x y : ↑(multiequalizer I)) (h : ∀ t : I.L, Multiequalizer.ι I t x = Multiequalizer.ι I t y) : x = y := by apply Concrete.limit_ext @@ -259,7 +259,7 @@ theorem multiequalizer_ext {I : MulticospanIndex.{w} C} [HasMultiequalizer I] simp [h] /-- An auxiliary equivalence to be used in `multiequalizerEquiv` below. -/ -def multiequalizerEquivAux (I : MulticospanIndex C) : +def multiequalizerEquivAux (I : MulticospanIndex.{w, w'} C) : (I.multicospan ⋙ forget C).sections ≃ { x : ∀ i : I.L, I.left i // ∀ i : I.R, I.fst i (x _) = I.snd i (x _) } where toFun x := @@ -292,17 +292,17 @@ def multiequalizerEquivAux (I : MulticospanIndex C) : /-- The equivalence between the noncomputable multiequalizer and the concrete multiequalizer. -/ -noncomputable def multiequalizerEquiv (I : MulticospanIndex.{w} C) [HasMultiequalizer I] +noncomputable def multiequalizerEquiv (I : MulticospanIndex.{w, w'} C) [HasMultiequalizer I] [PreservesLimit I.multicospan (forget C)] : (multiequalizer I : C) ≃ { x : ∀ i : I.L, I.left i // ∀ i : I.R, I.fst i (x _) = I.snd i (x _) } := letI h1 := limit.isLimit I.multicospan letI h2 := isLimitOfPreserves (forget C) h1 - letI E := h2.conePointUniqueUpToIso (Types.limitConeIsLimit.{w, v} _) - Equiv.trans E.toEquiv (Concrete.multiequalizerEquivAux.{w, v} I) + letI E := h2.conePointUniqueUpToIso (Types.limitConeIsLimit.{max w w', v} _) + Equiv.trans E.toEquiv (Concrete.multiequalizerEquivAux I) @[simp] -theorem multiequalizerEquiv_apply (I : MulticospanIndex.{w} C) [HasMultiequalizer I] +theorem multiequalizerEquiv_apply (I : MulticospanIndex.{w, w'} C) [HasMultiequalizer I] [PreservesLimit I.multicospan (forget C)] (x : ↑(multiequalizer I)) (i : I.L) : ((Concrete.multiequalizerEquiv I) x : ∀ i : I.L, I.left i) i = Multiequalizer.ι I i x := rfl diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Multiequalizer.lean b/Mathlib/CategoryTheory/Limits/Shapes/Multiequalizer.lean index d38347887cbc5..a8f12eea475ca 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Multiequalizer.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Multiequalizer.lean @@ -31,29 +31,29 @@ namespace CategoryTheory.Limits open CategoryTheory -universe w v u +universe w w' v u /-- The type underlying the multiequalizer diagram. -/ --@[nolint unused_arguments] -inductive WalkingMulticospan {L R : Type w} (fst snd : R → L) : Type w +inductive WalkingMulticospan {L : Type w} {R : Type w'} (fst snd : R → L) : Type max w w' | left : L → WalkingMulticospan fst snd | right : R → WalkingMulticospan fst snd /-- The type underlying the multiecoqualizer diagram. -/ --@[nolint unused_arguments] -inductive WalkingMultispan {L R : Type w} (fst snd : L → R) : Type w +inductive WalkingMultispan {L : Type w} {R : Type w'} (fst snd : L → R) : Type max w w' | left : L → WalkingMultispan fst snd | right : R → WalkingMultispan fst snd namespace WalkingMulticospan -variable {L R : Type w} {fst snd : R → L} +variable {L : Type w} {R : Type w'} {fst snd : R → L} instance [Inhabited L] : Inhabited (WalkingMulticospan fst snd) := ⟨left default⟩ /-- Morphisms for `WalkingMulticospan`. -/ -inductive Hom : ∀ _ _ : WalkingMulticospan fst snd, Type w +inductive Hom : ∀ _ _ : WalkingMulticospan fst snd, Type max w w' | id (A) : Hom A A | fst (b) : Hom (left (fst b)) (right b) | snd (b) : Hom (left (snd b)) (right b) @@ -94,13 +94,13 @@ end WalkingMulticospan namespace WalkingMultispan -variable {L R : Type v} {fst snd : L → R} +variable {L : Type w} {R : Type w'} {fst snd : L → R} instance [Inhabited L] : Inhabited (WalkingMultispan fst snd) := ⟨left default⟩ /-- Morphisms for `WalkingMultispan`. -/ -inductive Hom : ∀ _ _ : WalkingMultispan fst snd, Type v +inductive Hom : ∀ _ _ : WalkingMultispan fst snd, Type max w w' | id (A) : Hom A A | fst (a) : Hom (left a) (right (fst a)) | snd (a) : Hom (left a) (right (snd a)) @@ -139,10 +139,11 @@ lemma Hom.comp_eq_comp {X Y Z : WalkingMultispan fst snd} end WalkingMultispan /-- This is a structure encapsulating the data necessary to define a `Multicospan`. -/ --- Porting note(#5171): linter not ported yet --- @[nolint has_nonempty_instance] +-- Porting note(#5171): has_nonempty_instance linter not ported yet +@[nolint checkUnivs] structure MulticospanIndex (C : Type u) [Category.{v} C] where - (L R : Type w) + (L : Type w) + (R : Type w') (fstTo sndTo : R → L) left : L → C right : R → C @@ -150,10 +151,11 @@ structure MulticospanIndex (C : Type u) [Category.{v} C] where snd : ∀ b, left (sndTo b) ⟶ right b /-- This is a structure encapsulating the data necessary to define a `Multispan`. -/ --- Porting note(#5171): linter not ported yet --- @[nolint has_nonempty_instance] +-- Porting note(#5171): has_nonempty_instance linter not ported yet +@[nolint checkUnivs] structure MultispanIndex (C : Type u) [Category.{v} C] where - (L R : Type w) + (L : Type w) + (R : Type w') (fstFrom sndFrom : L → R) left : L → C right : R → C @@ -162,7 +164,7 @@ structure MultispanIndex (C : Type u) [Category.{v} C] where namespace MulticospanIndex -variable {C : Type u} [Category.{v} C] (I : MulticospanIndex.{w} C) +variable {C : Type u} [Category.{v} C] (I : MulticospanIndex.{w, w'} C) /-- The multicospan associated to `I : MulticospanIndex`. -/ @[simps] @@ -210,7 +212,7 @@ end MulticospanIndex namespace MultispanIndex -variable {C : Type u} [Category.{v} C] (I : MultispanIndex.{w} C) +variable {C : Type u} [Category.{v} C] (I : MultispanIndex.{w, w'} C) /-- The multispan associated to `I : MultispanIndex`. -/ def multispan : WalkingMultispan I.fstFrom I.sndFrom ⥤ C where @@ -276,18 +278,18 @@ variable {C : Type u} [Category.{v} C] /-- A multifork is a cone over a multicospan. -/ -- Porting note(#5171): linter not ported yet -- @[nolint has_nonempty_instance] -abbrev Multifork (I : MulticospanIndex.{w} C) := +abbrev Multifork (I : MulticospanIndex.{w, w'} C) := Cone I.multicospan /-- A multicofork is a cocone over a multispan. -/ -- Porting note(#5171): linter not ported yet -- @[nolint has_nonempty_instance] -abbrev Multicofork (I : MultispanIndex.{w} C) := +abbrev Multicofork (I : MultispanIndex.{w, w'} C) := Cocone I.multispan namespace Multifork -variable {I : MulticospanIndex.{w} C} (K : Multifork I) +variable {I : MulticospanIndex.{w, w'} C} (K : Multifork I) /-- The maps from the cone point of a multifork to the objects on the left. -/ def ι (a : I.L) : K.pt ⟶ I.left a := @@ -315,7 +317,7 @@ theorem hom_comp_ι (K₁ K₂ : Multifork I) (f : K₁ ⟶ K₂) (j : I.L) : f. /-- Construct a multifork using a collection `ι` of morphisms. -/ @[simps] -def ofι (I : MulticospanIndex.{w} C) (P : C) (ι : ∀ a, P ⟶ I.left a) +def ofι (I : MulticospanIndex.{w, w'} C) (P : C) (ι : ∀ a, P ⟶ I.left a) (w : ∀ b, ι (I.fstTo b) ≫ I.fst b = ι (I.sndTo b) ≫ I.snd b) : Multifork I where pt := P π := @@ -438,7 +440,7 @@ end Multifork namespace MulticospanIndex -variable (I : MulticospanIndex.{w} C) [HasProduct I.left] [HasProduct I.right] +variable (I : MulticospanIndex.{w, w'} C) [HasProduct I.left] [HasProduct I.right] --attribute [local tidy] tactic.case_bash @@ -485,7 +487,7 @@ end MulticospanIndex namespace Multicofork -variable {I : MultispanIndex.{w} C} (K : Multicofork I) +variable {I : MultispanIndex.{w, w'} C} (K : Multicofork I) /-- The maps to the cocone point of a multicofork from the objects on the right. -/ def π (b : I.R) : I.right b ⟶ K.pt := @@ -511,7 +513,7 @@ lemma π_comp_hom (K₁ K₂ : Multicofork I) (f : K₁ ⟶ K₂) (b : I.R) : K /-- Construct a multicofork using a collection `π` of morphisms. -/ @[simps] -def ofπ (I : MultispanIndex.{w} C) (P : C) (π : ∀ b, I.right b ⟶ P) +def ofπ (I : MultispanIndex.{w, w'} C) (P : C) (π : ∀ b, I.right b ⟶ P) (w : ∀ a, I.fst a ≫ π (I.fstFrom a) = I.snd a ≫ π (I.sndFrom a)) : Multicofork I where pt := P ι := @@ -614,7 +616,7 @@ end Multicofork namespace MultispanIndex -variable (I : MultispanIndex.{w} C) [HasCoproduct I.left] [HasCoproduct I.right] +variable (I : MultispanIndex.{w, w'} C) [HasCoproduct I.left] [HasCoproduct I.right] --attribute [local tidy] tactic.case_bash @@ -673,27 +675,27 @@ end MultispanIndex /-- For `I : MulticospanIndex C`, we say that it has a multiequalizer if the associated multicospan has a limit. -/ -abbrev HasMultiequalizer (I : MulticospanIndex.{w} C) := +abbrev HasMultiequalizer (I : MulticospanIndex.{w, w'} C) := HasLimit I.multicospan noncomputable section /-- The multiequalizer of `I : MulticospanIndex C`. -/ -abbrev multiequalizer (I : MulticospanIndex.{w} C) [HasMultiequalizer I] : C := +abbrev multiequalizer (I : MulticospanIndex.{w, w'} C) [HasMultiequalizer I] : C := limit I.multicospan /-- For `I : MultispanIndex C`, we say that it has a multicoequalizer if the associated multicospan has a limit. -/ -abbrev HasMulticoequalizer (I : MultispanIndex.{w} C) := +abbrev HasMulticoequalizer (I : MultispanIndex.{w, w'} C) := HasColimit I.multispan /-- The multiecoqualizer of `I : MultispanIndex C`. -/ -abbrev multicoequalizer (I : MultispanIndex.{w} C) [HasMulticoequalizer I] : C := +abbrev multicoequalizer (I : MultispanIndex.{w, w'} C) [HasMulticoequalizer I] : C := colimit I.multispan namespace Multiequalizer -variable (I : MulticospanIndex.{w} C) [HasMultiequalizer I] +variable (I : MulticospanIndex.{w, w'} C) [HasMultiequalizer I] /-- The canonical map from the multiequalizer to the objects on the left. -/ abbrev ι (a : I.L) : multiequalizer I ⟶ I.left a := @@ -758,7 +760,7 @@ end Multiequalizer namespace Multicoequalizer -variable (I : MultispanIndex.{w} C) [HasMulticoequalizer I] +variable (I : MultispanIndex.{w, w'} C) [HasMulticoequalizer I] /-- The canonical map from the multiequalizer to the objects on the left. -/ abbrev π (b : I.R) : I.right b ⟶ multicoequalizer I := From bc821d31df146cf3e18fbe53a94013c37a399dac Mon Sep 17 00:00:00 2001 From: "Tristan F.-R." Date: Thu, 24 Oct 2024 12:28:54 +0000 Subject: [PATCH 371/505] golf(Wiedijk/InverseTriangleSum): cleanup (#18152) After #18100, the `is_pos` and `is_neg` cases don't need to be separate any more, and `n` is not referenced explicitly. Co-authored-by: LeoDog896 --- Archive/Wiedijk100Theorems/InverseTriangleSum.lean | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/Archive/Wiedijk100Theorems/InverseTriangleSum.lean b/Archive/Wiedijk100Theorems/InverseTriangleSum.lean index e1ec0f91f2648..af36acfaa873d 100644 --- a/Archive/Wiedijk100Theorems/InverseTriangleSum.lean +++ b/Archive/Wiedijk100Theorems/InverseTriangleSum.lean @@ -27,8 +27,8 @@ open Finset /-- **Sum of the Reciprocals of the Triangular Numbers** -/ theorem Theorems100.inverse_triangle_sum : ∀ n, ∑ k ∈ range n, (2 : ℚ) / (k * (k + 1)) = if n = 0 then 0 else 2 - (2 : ℚ) / n := by - refine sum_range_induction _ _ (if_pos rfl) ?_ - rintro (_ | n) - · rw [if_neg, if_pos] <;> norm_num + refine sum_range_induction _ _ rfl ?_ + rintro (_ | _) + · norm_num field_simp ring From 688184279352ca105e037a7d657115bccbc4e079 Mon Sep 17 00:00:00 2001 From: damiano Date: Thu, 24 Oct 2024 13:13:03 +0000 Subject: [PATCH 372/505] fix: sign for percentage and preserve data file (#18170) Fixes the printing of a sign in the reported percentages. [Zulip](https://leanprover.zulipchat.com/#narrow/channel/113488-general/topic/bench.20summaries/near/478694897) --- scripts/bench_summary.lean | 10 +++++----- 1 file changed, 5 insertions(+), 5 deletions(-) diff --git a/scripts/bench_summary.lean b/scripts/bench_summary.lean index b7c521875b401..909dc9826ac4b 100644 --- a/scripts/bench_summary.lean +++ b/scripts/bench_summary.lean @@ -80,9 +80,7 @@ def formatFile (file : String) : String := s!"`{file.dropWhile (!·.isAlpha)}`" ``| `file` | ±x.y⬝10⁹ | ±z.w% |`` (technically, without the spaces). -/ def summary (bc : Bench) : String := - let reldiff := bc.reldiff - let (sgnf : Float) := if reldiff < 0 then -1 else 1 - let middle := [formatFile bc.file, formatDiff bc.diff, formatPercent (sgnf * reldiff)] + let middle := [formatFile bc.file, formatDiff bc.diff, formatPercent bc.reldiff] "|".intercalate (""::middle ++ [""]) /-- @@ -188,7 +186,7 @@ def addBenchSummaryComment (PR : Nat) (repo : String) (tempFile : String := "ben { cmd := "curl" args := #[s!"http://speed.lean-fro.org/mathlib4/api/compare/{source}/to/{target}?all_values=true"] } let bench ← IO.Process.run curlSpeedCenter - IO.FS.writeFile tempFile bench + IO.FS.writeFile (tempFile ++ ".src") bench -- Extract all instruction changes whose magnitude is larger than `threshold`. let threshold := s!"{10 ^ 9}" let jq1 : IO.Process.SpawnArgs := @@ -196,8 +194,10 @@ def addBenchSummaryComment (PR : Nat) (repo : String) (tempFile : String := "ben args := #["-r", "--arg", "thr", threshold, ".differences | .[] | ($thr|tonumber) as $th | select(.dimension.metric == \"instructions\" and ((.diff >= $th) or (.diff <= -$th)))", - tempFile] } + (tempFile ++ ".src")] } let firstFilter ← IO.Process.run jq1 + -- we leave `tempFile.src` unchanged and we switch to updating `tempfile`: this is useful for + -- debugging, as it preserves the original data downloaded from the speed-center IO.FS.writeFile tempFile firstFilter -- Write these in compact form, in the format suitable for `benchOutput`. let jq2 : IO.Process.SpawnArgs := From 3693a5a57e378f05c10214413ad7976295d762f7 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Thu, 24 Oct 2024 13:22:54 +0000 Subject: [PATCH 373/505] chore: split NumberTheory/NumberField/Discriminant (#18168) Splits NumberTheory/NumberField/Discriminant in Defs and Basic. Defs suffices for Cyclotomic/Discriminant, and doesn't require much analysis. Basic contains Hermite's theorem. --- Mathlib.lean | 3 +- .../NumberTheory/Cyclotomic/Discriminant.lean | 3 +- .../NumberTheory/NumberField/ClassNumber.lean | 3 +- .../Basic.lean} | 91 +------------- .../NumberField/Discriminant/Defs.lean | 119 ++++++++++++++++++ 5 files changed, 125 insertions(+), 94 deletions(-) rename Mathlib/NumberTheory/NumberField/{Discriminant.lean => Discriminant/Basic.lean} (83%) create mode 100644 Mathlib/NumberTheory/NumberField/Discriminant/Defs.lean diff --git a/Mathlib.lean b/Mathlib.lean index d3acec601583c..6e13fcffd54e8 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3594,7 +3594,8 @@ import Mathlib.NumberTheory.NumberField.CanonicalEmbedding.Basic import Mathlib.NumberTheory.NumberField.CanonicalEmbedding.ConvexBody import Mathlib.NumberTheory.NumberField.CanonicalEmbedding.FundamentalCone import Mathlib.NumberTheory.NumberField.ClassNumber -import Mathlib.NumberTheory.NumberField.Discriminant +import Mathlib.NumberTheory.NumberField.Discriminant.Basic +import Mathlib.NumberTheory.NumberField.Discriminant.Defs import Mathlib.NumberTheory.NumberField.Embeddings import Mathlib.NumberTheory.NumberField.EquivReindex import Mathlib.NumberTheory.NumberField.FractionalIdeal diff --git a/Mathlib/NumberTheory/Cyclotomic/Discriminant.lean b/Mathlib/NumberTheory/Cyclotomic/Discriminant.lean index c4b31c98a929a..0e61f7398af63 100644 --- a/Mathlib/NumberTheory/Cyclotomic/Discriminant.lean +++ b/Mathlib/NumberTheory/Cyclotomic/Discriminant.lean @@ -4,7 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Riccardo Brasca -/ import Mathlib.NumberTheory.Cyclotomic.PrimitiveRoots -import Mathlib.NumberTheory.NumberField.Discriminant +import Mathlib.RingTheory.DedekindDomain.Dvr +import Mathlib.NumberTheory.NumberField.Discriminant.Defs /-! # Discriminant of cyclotomic fields diff --git a/Mathlib/NumberTheory/NumberField/ClassNumber.lean b/Mathlib/NumberTheory/NumberField/ClassNumber.lean index 06ed53e2b5316..769d1f2df3687 100644 --- a/Mathlib/NumberTheory/NumberField/ClassNumber.lean +++ b/Mathlib/NumberTheory/NumberField/ClassNumber.lean @@ -5,7 +5,7 @@ Authors: Anne Baanen -/ import Mathlib.NumberTheory.ClassNumber.AdmissibleAbs import Mathlib.NumberTheory.ClassNumber.Finite -import Mathlib.NumberTheory.NumberField.Discriminant +import Mathlib.NumberTheory.NumberField.Discriminant.Basic /-! # Class numbers of number fields @@ -19,7 +19,6 @@ on the class number. cardinality of the class group of its ring of integers -/ - namespace NumberField variable (K : Type*) [Field K] [NumberField K] diff --git a/Mathlib/NumberTheory/NumberField/Discriminant.lean b/Mathlib/NumberTheory/NumberField/Discriminant/Basic.lean similarity index 83% rename from Mathlib/NumberTheory/NumberField/Discriminant.lean rename to Mathlib/NumberTheory/NumberField/Discriminant/Basic.lean index 393c5e32016a0..26397e665bada 100644 --- a/Mathlib/NumberTheory/NumberField/Discriminant.lean +++ b/Mathlib/NumberTheory/NumberField/Discriminant/Basic.lean @@ -6,15 +6,12 @@ Authors: Xavier Roblot import Mathlib.Data.Real.Pi.Bounds import Mathlib.NumberTheory.NumberField.CanonicalEmbedding.ConvexBody import Mathlib.Tactic.Rify +import Mathlib.NumberTheory.NumberField.Discriminant.Defs /-! # Number field discriminant This file defines the discriminant of a number field. -## Main definitions - -* `NumberField.discr`: the absolute discriminant of a number field. - ## Main result * `NumberField.abs_discr_gt_two`: **Hermite-Minkowski Theorem**. A nontrivial number field has @@ -38,34 +35,6 @@ open scoped Classical Real nonZeroDivisors variable (K : Type*) [Field K] [NumberField K] -/-- The absolute discriminant of a number field. -/ -noncomputable abbrev discr : ℤ := Algebra.discr ℤ (RingOfIntegers.basis K) - -theorem coe_discr : (discr K : ℚ) = Algebra.discr ℚ (integralBasis K) := - (Algebra.discr_localizationLocalization ℤ _ K (RingOfIntegers.basis K)).symm - -theorem discr_ne_zero : discr K ≠ 0 := by - rw [← (Int.cast_injective (α := ℚ)).ne_iff, coe_discr] - exact Algebra.discr_not_zero_of_basis ℚ (integralBasis K) - -theorem discr_eq_discr {ι : Type*} [Fintype ι] [DecidableEq ι] (b : Basis ι ℤ (𝓞 K)) : - Algebra.discr ℤ b = discr K := by - let b₀ := Basis.reindex (RingOfIntegers.basis K) (Basis.indexEquiv (RingOfIntegers.basis K) b) - rw [Algebra.discr_eq_discr (𝓞 K) b b₀, Basis.coe_reindex, Algebra.discr_reindex] - -theorem discr_eq_discr_of_algEquiv {L : Type*} [Field L] [NumberField L] (f : K ≃ₐ[ℚ] L) : - discr K = discr L := by - let f₀ : 𝓞 K ≃ₗ[ℤ] 𝓞 L := (f.restrictScalars ℤ).mapIntegralClosure.toLinearEquiv - rw [← Rat.intCast_inj, coe_discr, Algebra.discr_eq_discr_of_algEquiv (integralBasis K) f, - ← discr_eq_discr L ((RingOfIntegers.basis K).map f₀)] - change _ = algebraMap ℤ ℚ _ - rw [← Algebra.discr_localizationLocalization ℤ (nonZeroDivisors ℤ) L] - congr - ext - simp only [Function.comp_apply, integralBasis_apply, Basis.localizationLocalization_apply, - Basis.map_apply] - rfl - open MeasureTheory MeasureTheory.Measure ZSpan NumberField.mixedEmbedding NumberField.InfinitePlace ENNReal NNReal Complex @@ -441,61 +410,3 @@ theorem _root_.NumberField.finite_of_discr_bdd : end hermiteTheorem end NumberField - -namespace Rat - -open NumberField - -/-- The absolute discriminant of the number field `ℚ` is 1. -/ -@[simp] -theorem numberField_discr : discr ℚ = 1 := by - let b : Basis (Fin 1) ℤ (𝓞 ℚ) := - Basis.map (Basis.singleton (Fin 1) ℤ) ringOfIntegersEquiv.toAddEquiv.toIntLinearEquiv.symm - calc NumberField.discr ℚ - _ = Algebra.discr ℤ b := by convert (discr_eq_discr ℚ b).symm - _ = Algebra.trace ℤ (𝓞 ℚ) (b default * b default) := by - rw [Algebra.discr_def, Matrix.det_unique, Algebra.traceMatrix_apply, Algebra.traceForm_apply] - _ = Algebra.trace ℤ (𝓞 ℚ) 1 := by - rw [Basis.map_apply, RingEquiv.toAddEquiv_eq_coe, AddEquiv.toIntLinearEquiv_symm, - AddEquiv.coe_toIntLinearEquiv, Basis.singleton_apply, - show (AddEquiv.symm ↑ringOfIntegersEquiv) (1 : ℤ) = ringOfIntegersEquiv.symm 1 by rfl, - map_one, mul_one] - _ = 1 := by rw [Algebra.trace_eq_matrix_trace b]; norm_num - -alias _root_.NumberField.discr_rat := numberField_discr - -end Rat - -variable {ι ι'} (K) [Field K] [DecidableEq ι] [DecidableEq ι'] [Fintype ι] [Fintype ι'] - -/-- If `b` and `b'` are `ℚ`-bases of a number field `K` such that -`∀ i j, IsIntegral ℤ (b.toMatrix b' i j)` and `∀ i j, IsIntegral ℤ (b'.toMatrix b i j)` then -`discr ℚ b = discr ℚ b'`. -/ -theorem Algebra.discr_eq_discr_of_toMatrix_coeff_isIntegral [NumberField K] - {b : Basis ι ℚ K} {b' : Basis ι' ℚ K} (h : ∀ i j, IsIntegral ℤ (b.toMatrix b' i j)) - (h' : ∀ i j, IsIntegral ℤ (b'.toMatrix b i j)) : discr ℚ b = discr ℚ b' := by - replace h' : ∀ i j, IsIntegral ℤ (b'.toMatrix (b.reindex (b.indexEquiv b')) i j) := by - intro i j - convert h' i ((b.indexEquiv b').symm j) - simp [Basis.toMatrix_apply] - classical - rw [← (b.reindex (b.indexEquiv b')).toMatrix_map_vecMul b', discr_of_matrix_vecMul, - ← one_mul (discr ℚ b), Basis.coe_reindex, discr_reindex] - congr - have hint : IsIntegral ℤ ((b.reindex (b.indexEquiv b')).toMatrix b').det := - IsIntegral.det fun i j => h _ _ - obtain ⟨r, hr⟩ := IsIntegrallyClosed.isIntegral_iff.1 hint - have hunit : IsUnit r := by - have : IsIntegral ℤ (b'.toMatrix (b.reindex (b.indexEquiv b'))).det := - IsIntegral.det fun i j => h' _ _ - obtain ⟨r', hr'⟩ := IsIntegrallyClosed.isIntegral_iff.1 this - refine isUnit_iff_exists_inv.2 ⟨r', ?_⟩ - suffices algebraMap ℤ ℚ (r * r') = 1 by - rw [← RingHom.map_one (algebraMap ℤ ℚ)] at this - exact (IsFractionRing.injective ℤ ℚ) this - rw [RingHom.map_mul, hr, hr', ← Matrix.det_mul, - Basis.toMatrix_mul_toMatrix_flip, Matrix.det_one] - rw [← RingHom.map_one (algebraMap ℤ ℚ), ← hr] - cases' Int.isUnit_iff.1 hunit with hp hm - · simp [hp] - · simp [hm] diff --git a/Mathlib/NumberTheory/NumberField/Discriminant/Defs.lean b/Mathlib/NumberTheory/NumberField/Discriminant/Defs.lean new file mode 100644 index 0000000000000..f9ccf3c816e85 --- /dev/null +++ b/Mathlib/NumberTheory/NumberField/Discriminant/Defs.lean @@ -0,0 +1,119 @@ +/- +Copyright (c) 2023 Xavier Roblot. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Xavier Roblot +-/ +import Init.Data.ULift +import Init.Data.Fin.Fold +import Init.Data.List.Nat.Pairwise +import Init.Data.List.Nat.Range +import Mathlib.NumberTheory.NumberField.Basic +import Mathlib.RingTheory.Localization.NormTrace +import Mathlib.Analysis.Normed.Field.Lemmas + +/-! +# Number field discriminant +This file defines the discriminant of a number field. + +## Main definitions + +* `NumberField.discr`: the absolute discriminant of a number field. + +## Tags +number field, discriminant +-/ + +-- TODO. Rewrite some of the FLT results on the disciminant using the definitions and results of +-- this file + +namespace NumberField + +variable (K : Type*) [Field K] [NumberField K] + +/-- The absolute discriminant of a number field. -/ +noncomputable abbrev discr : ℤ := Algebra.discr ℤ (RingOfIntegers.basis K) + +theorem coe_discr : (discr K : ℚ) = Algebra.discr ℚ (integralBasis K) := + (Algebra.discr_localizationLocalization ℤ _ K (RingOfIntegers.basis K)).symm + +theorem discr_ne_zero : discr K ≠ 0 := by + rw [← (Int.cast_injective (α := ℚ)).ne_iff, coe_discr] + exact Algebra.discr_not_zero_of_basis ℚ (integralBasis K) + +theorem discr_eq_discr {ι : Type*} [Fintype ι] [DecidableEq ι] (b : Basis ι ℤ (𝓞 K)) : + Algebra.discr ℤ b = discr K := by + let b₀ := Basis.reindex (RingOfIntegers.basis K) (Basis.indexEquiv (RingOfIntegers.basis K) b) + rw [Algebra.discr_eq_discr (𝓞 K) b b₀, Basis.coe_reindex, Algebra.discr_reindex] + +theorem discr_eq_discr_of_algEquiv {L : Type*} [Field L] [NumberField L] (f : K ≃ₐ[ℚ] L) : + discr K = discr L := by + let f₀ : 𝓞 K ≃ₗ[ℤ] 𝓞 L := (f.restrictScalars ℤ).mapIntegralClosure.toLinearEquiv + rw [← Rat.intCast_inj, coe_discr, Algebra.discr_eq_discr_of_algEquiv (integralBasis K) f, + ← discr_eq_discr L ((RingOfIntegers.basis K).map f₀)] + change _ = algebraMap ℤ ℚ _ + rw [← Algebra.discr_localizationLocalization ℤ (nonZeroDivisors ℤ) L] + congr + ext + simp only [Function.comp_apply, integralBasis_apply, Basis.localizationLocalization_apply, + Basis.map_apply] + rfl + +end NumberField + +namespace Rat + +open NumberField + +/-- The absolute discriminant of the number field `ℚ` is 1. -/ +@[simp] +theorem numberField_discr : discr ℚ = 1 := by + let b : Basis (Fin 1) ℤ (𝓞 ℚ) := + Basis.map (Basis.singleton (Fin 1) ℤ) ringOfIntegersEquiv.toAddEquiv.toIntLinearEquiv.symm + calc NumberField.discr ℚ + _ = Algebra.discr ℤ b := by convert (discr_eq_discr ℚ b).symm + _ = Algebra.trace ℤ (𝓞 ℚ) (b default * b default) := by + rw [Algebra.discr_def, Matrix.det_unique, Algebra.traceMatrix_apply, Algebra.traceForm_apply] + _ = Algebra.trace ℤ (𝓞 ℚ) 1 := by + rw [Basis.map_apply, RingEquiv.toAddEquiv_eq_coe, AddEquiv.toIntLinearEquiv_symm, + AddEquiv.coe_toIntLinearEquiv, Basis.singleton_apply, + show (AddEquiv.symm ↑ringOfIntegersEquiv) (1 : ℤ) = ringOfIntegersEquiv.symm 1 by rfl, + map_one, mul_one] + _ = 1 := by rw [Algebra.trace_eq_matrix_trace b]; norm_num + +alias _root_.NumberField.discr_rat := numberField_discr + +end Rat + +variable {ι ι'} (K) [Field K] [DecidableEq ι] [DecidableEq ι'] [Fintype ι] [Fintype ι'] + +/-- If `b` and `b'` are `ℚ`-bases of a number field `K` such that +`∀ i j, IsIntegral ℤ (b.toMatrix b' i j)` and `∀ i j, IsIntegral ℤ (b'.toMatrix b i j)` then +`discr ℚ b = discr ℚ b'`. -/ +theorem Algebra.discr_eq_discr_of_toMatrix_coeff_isIntegral [NumberField K] + {b : Basis ι ℚ K} {b' : Basis ι' ℚ K} (h : ∀ i j, IsIntegral ℤ (b.toMatrix b' i j)) + (h' : ∀ i j, IsIntegral ℤ (b'.toMatrix b i j)) : discr ℚ b = discr ℚ b' := by + replace h' : ∀ i j, IsIntegral ℤ (b'.toMatrix (b.reindex (b.indexEquiv b')) i j) := by + intro i j + convert h' i ((b.indexEquiv b').symm j) + simp [Basis.toMatrix_apply] + classical + rw [← (b.reindex (b.indexEquiv b')).toMatrix_map_vecMul b', discr_of_matrix_vecMul, + ← one_mul (discr ℚ b), Basis.coe_reindex, discr_reindex] + congr + have hint : IsIntegral ℤ ((b.reindex (b.indexEquiv b')).toMatrix b').det := + IsIntegral.det fun i j => h _ _ + obtain ⟨r, hr⟩ := IsIntegrallyClosed.isIntegral_iff.1 hint + have hunit : IsUnit r := by + have : IsIntegral ℤ (b'.toMatrix (b.reindex (b.indexEquiv b'))).det := + IsIntegral.det fun i j => h' _ _ + obtain ⟨r', hr'⟩ := IsIntegrallyClosed.isIntegral_iff.1 this + refine isUnit_iff_exists_inv.2 ⟨r', ?_⟩ + suffices algebraMap ℤ ℚ (r * r') = 1 by + rw [← RingHom.map_one (algebraMap ℤ ℚ)] at this + exact (IsFractionRing.injective ℤ ℚ) this + rw [RingHom.map_mul, hr, hr', ← Matrix.det_mul, + Basis.toMatrix_mul_toMatrix_flip, Matrix.det_one] + rw [← RingHom.map_one (algebraMap ℤ ℚ), ← hr] + cases' Int.isUnit_iff.1 hunit with hp hm + · simp [hp] + · simp [hm] From f72cbf5998335b7b2863a003a2ce58c11e982418 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Thu, 24 Oct 2024 14:10:01 +0000 Subject: [PATCH 374/505] feat(AlgebraicGeometry): define immersions (#18154) Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- Mathlib.lean | 1 + .../Morphisms/ClosedImmersion.lean | 22 +-- .../Morphisms/Immersion.lean | 167 ++++++++++++++++++ 3 files changed, 180 insertions(+), 10 deletions(-) create mode 100644 Mathlib/AlgebraicGeometry/Morphisms/Immersion.lean diff --git a/Mathlib.lean b/Mathlib.lean index 6e13fcffd54e8..fb22c74a4bce0 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -880,6 +880,7 @@ import Mathlib.AlgebraicGeometry.Morphisms.ClosedImmersion import Mathlib.AlgebraicGeometry.Morphisms.Constructors import Mathlib.AlgebraicGeometry.Morphisms.FinitePresentation import Mathlib.AlgebraicGeometry.Morphisms.FiniteType +import Mathlib.AlgebraicGeometry.Morphisms.Immersion import Mathlib.AlgebraicGeometry.Morphisms.IsIso import Mathlib.AlgebraicGeometry.Morphisms.OpenImmersion import Mathlib.AlgebraicGeometry.Morphisms.Preimmersion diff --git a/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean b/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean index e6170d7e2cc33..c95fb9628853a 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean @@ -43,12 +43,14 @@ class IsClosedImmersion {X Y : Scheme} (f : X ⟶ Y) : Prop where base_closed : IsClosedEmbedding f.base surj_on_stalks : ∀ x, Function.Surjective (f.stalkMap x) -namespace IsClosedImmersion - -lemma isClosedEmbedding {X Y : Scheme} (f : X ⟶ Y) +lemma Scheme.Hom.isClosedEmbedding {X Y : Scheme} (f : X.Hom Y) [IsClosedImmersion f] : IsClosedEmbedding f.base := IsClosedImmersion.base_closed +namespace IsClosedImmersion + +@[deprecated (since := "2024-10-24")] +alias isClosedEmbedding := Scheme.Hom.isClosedEmbedding @[deprecated (since := "2024-10-20")] alias closedEmbedding := isClosedEmbedding @@ -128,11 +130,11 @@ where `g` is only required to be separated. theorem of_comp_isClosedImmersion {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) [IsClosedImmersion g] [IsClosedImmersion (f ≫ g)] : IsClosedImmersion f where base_closed := by - have h := isClosedEmbedding (f ≫ g) + have h := (f ≫ g).isClosedEmbedding simp only [Scheme.comp_coeBase, TopCat.coe_comp] at h refine .of_continuous_injective_isClosedMap (Scheme.Hom.continuous f) h.inj.of_comp ?_ intro Z hZ - rw [IsClosedEmbedding.closed_iff_image_closed (isClosedEmbedding g), + rw [IsClosedEmbedding.closed_iff_image_closed g.isClosedEmbedding, ← Set.image_comp] exact h.isClosedMap _ hZ surj_on_stalks x := by @@ -228,14 +230,14 @@ namespace IsClosedImmersion sections is injective, `f` is an isomorphism. -/ theorem isIso_of_injective_of_isAffine [IsClosedImmersion f] (hf : Function.Injective (f.app ⊤)) : IsIso f := (isIso_iff_stalk_iso f).mpr <| - have : CompactSpace X := (isClosedEmbedding f).compactSpace + have : CompactSpace X := f.isClosedEmbedding.compactSpace have hiso : IsIso f.base := TopCat.isIso_of_bijective_of_isClosedMap _ - ⟨(isClosedEmbedding f).inj, - surjective_of_isClosed_range_of_injective ((isClosedEmbedding f).isClosed_range) hf⟩ - ((isClosedEmbedding f).isClosedMap) + ⟨f.isClosedEmbedding.inj, + surjective_of_isClosed_range_of_injective f.isClosedEmbedding.isClosed_range hf⟩ + (f.isClosedEmbedding.isClosedMap) ⟨hiso, fun x ↦ (ConcreteCategory.isIso_iff_bijective _).mpr ⟨stalkMap_injective_of_isOpenMap_of_injective ((TopCat.homeoOfIso (asIso f.base)).isOpenMap) - (isClosedEmbedding f).inj hf _, f.stalkMap_surjective x⟩⟩ + f.isClosedEmbedding.inj hf _, f.stalkMap_surjective x⟩⟩ variable (f) diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Immersion.lean b/Mathlib/AlgebraicGeometry/Morphisms/Immersion.lean new file mode 100644 index 0000000000000..3dde16b8057b3 --- /dev/null +++ b/Mathlib/AlgebraicGeometry/Morphisms/Immersion.lean @@ -0,0 +1,167 @@ +/- +Copyright (c) 2024 Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Andrew Yang +-/ +import Mathlib.AlgebraicGeometry.Morphisms.Preimmersion +import Mathlib.AlgebraicGeometry.Morphisms.ClosedImmersion + +/-! + +# Immersions of schemes + +A morphism of schemes `f : X ⟶ Y` is an immersion if the underlying map of topological spaces +is a locally closed embedding, and the induced morphisms of stalks are all surjective. This is true +if and only if it can be factored into a closed immersion followed by an open immersion. + +# Main result +- `isImmersion_iff_exists`: + A morphism is a (locally-closed) immersion if and only if it can be factored into + a closed immersion followed by a (dominant) open immersion. + + +## TODO + +* Show that diagonal morphisms are immersions + +-/ + +universe v u + +open CategoryTheory + +namespace AlgebraicGeometry + +variable {X Y : Scheme.{u}} (f : X ⟶ Y) + +/-- A morphism of schemes `f : X ⟶ Y` is an immersion if +1. the underlying map of topological spaces is an embedding +2. the range of the map is locally closed +3. the induced morphisms of stalks are all surjective. -/ +@[mk_iff] +class IsImmersion (f : X ⟶ Y) extends IsPreimmersion f : Prop where + isLocallyClosed_range : IsLocallyClosed (Set.range f.base) + +lemma Scheme.Hom.isLocallyClosed_range (f : X.Hom Y) [IsImmersion f] : + IsLocallyClosed (Set.range f.base) := + IsImmersion.isLocallyClosed_range + +/-- +Given an immersion `f : X ⟶ Y`, this is the biggest open set `U ⊆ Y` containing the image of `X` +such that `X` is closed in `U`. +-/ +def Scheme.Hom.coborderRange (f : X.Hom Y) [IsImmersion f] : Y.Opens := + ⟨coborder (Set.range f.base), f.isLocallyClosed_range.isOpen_coborder⟩ + +/-- +The first part of the factorization of an immersion `f : X ⟶ Y` to a closed immersion +`f.liftCoborder : X ⟶ f.coborderRange` and a dominant open immersion `f.coborderRange.ι`. +-/ +noncomputable +def Scheme.Hom.liftCoborder (f : X.Hom Y) [IsImmersion f] : X ⟶ f.coborderRange := + IsOpenImmersion.lift f.coborderRange.ι f (by simpa using subset_coborder) + +/-- +Any (locally-closed) immersion can be factored into +a closed immersion followed by a (dominant) open immersion. +-/ +@[reassoc (attr := simp)] +lemma Scheme.Hom.liftCoborder_ι (f : X.Hom Y) [IsImmersion f] : + f.liftCoborder ≫ f.coborderRange.ι = f := + IsOpenImmersion.lift_fac _ _ _ + +instance [IsImmersion f] : IsClosedImmersion f.liftCoborder := by + have : IsPreimmersion (f.liftCoborder ≫ f.coborderRange.ι) := by + simp only [Scheme.Hom.liftCoborder_ι]; infer_instance + have : IsPreimmersion f.liftCoborder := .of_comp f.liftCoborder f.coborderRange.ι + refine .of_isPreimmersion _ ?_ + convert isClosed_preimage_val_coborder + apply Set.image_injective.mpr f.coborderRange.ι.embedding.inj + rw [← Set.range_comp, ← TopCat.coe_comp, ← Scheme.comp_base, f.liftCoborder_ι] + exact (Set.image_preimage_eq_of_subset (by simpa using subset_coborder)).symm + +instance [IsImmersion f] : IsDominant f.coborderRange.ι := by + rw [isDominant_iff, DenseRange, Scheme.Opens.range_ι] + exact dense_coborder + +lemma isImmersion_eq_inf : @IsImmersion = (@IsPreimmersion ⊓ + topologically fun {X Y} _ _ f ↦ IsLocallyClosed (Set.range f) : MorphismProperty Scheme) := by + ext; exact isImmersion_iff _ + +namespace IsImmersion + +instance : IsLocalAtTarget @IsImmersion := by + suffices IsLocalAtTarget (topologically fun {X Y} _ _ f ↦ IsLocallyClosed (Set.range f)) from + isImmersion_eq_inf ▸ inferInstance + apply (config := { allowSynthFailures := true }) topologically_isLocalAtTarget' + · refine { precomp := ?_, postcomp := ?_ } + · intro X Y Z i hi f hf + replace hi : IsIso i := hi + show IsLocallyClosed _ + simpa only [Scheme.comp_coeBase, TopCat.coe_comp, Set.range_comp, + Set.range_iff_surjective.mpr i.surjective, Set.image_univ] + · intro X Y Z i hi f hf + replace hi : IsIso i := hi + show IsLocallyClosed _ + simp only [Scheme.comp_coeBase, TopCat.coe_comp, Set.range_comp] + refine hf.image i.homeomorph.inducing ?_ + rw [Set.range_iff_surjective.mpr i.surjective] + exact isOpen_univ.isLocallyClosed + · simp_rw [Set.range_restrictPreimage] + exact fun _ _ _ e _ ↦ isLocallyClosed_iff_coe_preimage_of_iSup_eq_top e _ + +instance (priority := 900) {X Y : Scheme} (f : X ⟶ Y) [IsOpenImmersion f] : IsImmersion f where + isLocallyClosed_range := f.isOpenEmbedding.2.isLocallyClosed + +instance (priority := 900) {X Y : Scheme} (f : X ⟶ Y) [IsClosedImmersion f] : IsImmersion f where + isLocallyClosed_range := f.isClosedEmbedding.2.isLocallyClosed + +instance : MorphismProperty.IsMultiplicative @IsImmersion where + id_mem _ := inferInstance + comp_mem {X Y Z} f g hf hg := by + refine { __ := inferInstanceAs (IsPreimmersion (f ≫ g)), isLocallyClosed_range := ?_ } + simp only [Scheme.comp_coeBase, TopCat.coe_comp, Set.range_comp] + exact f.isLocallyClosed_range.image g.embedding.toInducing g.isLocallyClosed_range + +instance comp {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) [IsImmersion f] + [IsImmersion g] : IsImmersion (f ≫ g) := + MorphismProperty.IsStableUnderComposition.comp_mem f g inferInstance inferInstance + +variable {f} in +/-- +A morphism is a (locally-closed) immersion if and only if it can be factored into +a closed immersion followed by an open immersion. +-/ +lemma isImmersion_iff_exists : IsImmersion f ↔ ∃ (Z : Scheme) (g₁ : X ⟶ Z) (g₂ : Z ⟶ Y), + IsClosedImmersion g₁ ∧ IsOpenImmersion g₂ ∧ g₁ ≫ g₂ = f := + ⟨fun _ ↦ ⟨_, f.liftCoborder, f.coborderRange.ι, inferInstance, inferInstance, f.liftCoborder_ι⟩, + fun ⟨_, _, _, _, _, e⟩ ↦ e ▸ inferInstance⟩ + +theorem of_comp {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) [IsImmersion g] + [IsImmersion (f ≫ g)] : IsImmersion f where + __ := IsPreimmersion.of_comp f g + isLocallyClosed_range := by + rw [← Set.preimage_image_eq (Set.range _) g.embedding.inj] + have := (f ≫ g).isLocallyClosed_range.preimage g.base.2 + simpa only [Scheme.comp_coeBase, TopCat.coe_comp, Set.range_comp] using this + +theorem comp_iff {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) [IsImmersion g] : + IsImmersion (f ≫ g) ↔ IsImmersion f := + ⟨fun _ ↦ of_comp f g, fun _ ↦ inferInstance⟩ + +lemma stableUnderBaseChange : MorphismProperty.StableUnderBaseChange @IsImmersion := by + intros X Y Y' S f g f' g' H hg + let Z := Limits.pullback f g.coborderRange.ι + let e : Y' ⟶ Z := Limits.pullback.lift g' (f' ≫ g.liftCoborder) (by simpa using H.w.symm) + have : IsClosedImmersion e := by + have := (IsPullback.paste_horiz_iff (.of_hasPullback f g.coborderRange.ι) + (show e ≫ Limits.pullback.snd _ _ = _ from Limits.pullback.lift_snd _ _ _)).mp ?_ + · exact IsClosedImmersion.stableUnderBaseChange this.flip inferInstance + · simpa [e] using H.flip + rw [← Limits.pullback.lift_fst (f := f) (g := g.coborderRange.ι) g' (f' ≫ g.liftCoborder) + (by simpa using H.w.symm)] + infer_instance + +end IsImmersion + +end AlgebraicGeometry From 03be989844769da3c80cae464d4ff2878b0494e7 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Thu, 24 Oct 2024 15:01:54 +0000 Subject: [PATCH 375/505] chore(Algebra/Module): split `Defs` a bit further (#18169) There are quite a few imports in `Module/Defs.lean` that are not needed for defining modules. So this PR rearranges the results in the file to postpone those imports. Changes: * The result involving `Invertible` go to `Module/Basic.lean` * The results involving `AddMonoid.End` goes to the new file `Module/End.lean` * The results involving `Opposite` go to the new file `Module/Opposite.lean`; the previous `Module/Opposites.lean` is now `Module/Equiv/Opposite.lean` (note: now in singular to match the other `Opposite` file) --- Mathlib.lean | 4 +- Mathlib/Algebra/Algebra/Operations.lean | 2 +- Mathlib/Algebra/Algebra/Opposite.lean | 2 +- Mathlib/Algebra/Algebra/Rat.lean | 1 + .../Algebra/Group/Submonoid/Pointwise.lean | 1 + Mathlib/Algebra/Module/Basic.lean | 6 + Mathlib/Algebra/Module/Defs.lean | 101 +-------------- Mathlib/Algebra/Module/End.lean | 117 ++++++++++++++++++ Mathlib/Algebra/Module/Equiv/Basic.lean | 1 + .../{Opposites.lean => Equiv/Opposite.lean} | 24 +++- Mathlib/Algebra/Module/Hom.lean | 3 +- Mathlib/Algebra/Module/LinearMap/Basic.lean | 1 + Mathlib/Algebra/Module/LinearMap/Defs.lean | 13 +- Mathlib/Algebra/Module/LinearMap/End.lean | 1 + Mathlib/Algebra/Module/Opposite.lean | 43 +++++++ Mathlib/Algebra/NoZeroSMulDivisors/Basic.lean | 3 +- Mathlib/Algebra/Periodic.lean | 3 +- Mathlib/Algebra/Ring/AddAut.lean | 3 +- .../Algebra/Ring/Subsemiring/MulOpposite.lean | 1 + Mathlib/CategoryTheory/Linear/Basic.lean | 1 + Mathlib/CategoryTheory/Preadditive/Basic.lean | 2 +- Mathlib/Data/Set/Pointwise/SMul.lean | 1 + Mathlib/Data/ZMod/Basic.lean | 1 + .../LinearAlgebra/AffineSpace/Midpoint.lean | 1 + .../CliffordAlgebra/Conjugation.lean | 2 +- Mathlib/RingTheory/OreLocalization/Ring.lean | 3 +- Mathlib/RingTheory/TwoSidedIdeal/Basic.lean | 1 + Mathlib/Topology/Algebra/Module/Basic.lean | 12 +- .../Topology/Algebra/UniformMulAction.lean | 1 + scripts/noshake.json | 1 + 30 files changed, 225 insertions(+), 131 deletions(-) create mode 100644 Mathlib/Algebra/Module/End.lean rename Mathlib/Algebra/Module/{Opposites.lean => Equiv/Opposite.lean} (73%) create mode 100644 Mathlib/Algebra/Module/Opposite.lean diff --git a/Mathlib.lean b/Mathlib.lean index fb22c74a4bce0..2d79a052f5058 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -478,8 +478,10 @@ import Mathlib.Algebra.Module.Card import Mathlib.Algebra.Module.CharacterModule import Mathlib.Algebra.Module.DedekindDomain import Mathlib.Algebra.Module.Defs +import Mathlib.Algebra.Module.End import Mathlib.Algebra.Module.Equiv.Basic import Mathlib.Algebra.Module.Equiv.Defs +import Mathlib.Algebra.Module.Equiv.Opposite import Mathlib.Algebra.Module.FinitePresentation import Mathlib.Algebra.Module.GradedModule import Mathlib.Algebra.Module.Hom @@ -494,7 +496,7 @@ import Mathlib.Algebra.Module.LinearMap.Star import Mathlib.Algebra.Module.LocalizedModule import Mathlib.Algebra.Module.LocalizedModuleIntegers import Mathlib.Algebra.Module.MinimalAxioms -import Mathlib.Algebra.Module.Opposites +import Mathlib.Algebra.Module.Opposite import Mathlib.Algebra.Module.PID import Mathlib.Algebra.Module.Pi import Mathlib.Algebra.Module.PointwisePi diff --git a/Mathlib/Algebra/Algebra/Operations.lean b/Mathlib/Algebra/Algebra/Operations.lean index ed2b456a4db5b..ec2010316a560 100644 --- a/Mathlib/Algebra/Algebra/Operations.lean +++ b/Mathlib/Algebra/Algebra/Operations.lean @@ -8,7 +8,7 @@ import Mathlib.Algebra.Algebra.Equiv import Mathlib.Algebra.Algebra.Opposite import Mathlib.Algebra.Group.Pointwise.Finset.Basic import Mathlib.Algebra.GroupWithZero.NonZeroDivisors -import Mathlib.Algebra.Module.Opposites +import Mathlib.Algebra.Module.Opposite import Mathlib.Algebra.Module.Submodule.Bilinear import Mathlib.Algebra.Module.Submodule.Pointwise import Mathlib.Algebra.Order.Kleene diff --git a/Mathlib/Algebra/Algebra/Opposite.lean b/Mathlib/Algebra/Algebra/Opposite.lean index 11e27a60fe78f..f98ba51332141 100644 --- a/Mathlib/Algebra/Algebra/Opposite.lean +++ b/Mathlib/Algebra/Algebra/Opposite.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Eric Wieser -/ import Mathlib.Algebra.Algebra.Equiv -import Mathlib.Algebra.Module.Opposites +import Mathlib.Algebra.Module.Opposite import Mathlib.Algebra.Ring.Opposite /-! diff --git a/Mathlib/Algebra/Algebra/Rat.lean b/Mathlib/Algebra/Algebra/Rat.lean index ed05e1a90ee28..bfccd9246f030 100644 --- a/Mathlib/Algebra/Algebra/Rat.lean +++ b/Mathlib/Algebra/Algebra/Rat.lean @@ -5,6 +5,7 @@ Authors: Kenny Lau, Yury Kudryashov -/ import Mathlib.Algebra.Algebra.Defs import Mathlib.Algebra.GroupWithZero.Action.Basic +import Mathlib.Algebra.Module.End import Mathlib.Data.Rat.Cast.CharZero /-! diff --git a/Mathlib/Algebra/Group/Submonoid/Pointwise.lean b/Mathlib/Algebra/Group/Submonoid/Pointwise.lean index 9983f082743b8..fea277c83d7d4 100644 --- a/Mathlib/Algebra/Group/Submonoid/Pointwise.lean +++ b/Mathlib/Algebra/Group/Submonoid/Pointwise.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Eric Wieser. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Eric Wieser -/ +import Mathlib.Algebra.Group.Hom.End import Mathlib.Algebra.Group.Submonoid.Membership import Mathlib.Algebra.Order.BigOperators.Group.List import Mathlib.Data.Set.Pointwise.SMul diff --git a/Mathlib/Algebra/Module/Basic.lean b/Mathlib/Algebra/Module/Basic.lean index 18deb70317a22..0eeee9934e3d3 100644 --- a/Mathlib/Algebra/Module/Basic.lean +++ b/Mathlib/Algebra/Module/Basic.lean @@ -23,6 +23,12 @@ universe u v variable {α R M M₂ : Type*} +@[simp] +theorem invOf_two_smul_add_invOf_two_smul (R) [Semiring R] [AddCommMonoid M] [Module R M] + [Invertible (2 : R)] (x : M) : + (⅟ 2 : R) • x + (⅟ 2 : R) • x = x := + Convex.combo_self invOf_two_add_invOf_two _ + @[deprecated (since := "2024-04-17")] alias map_nat_cast_smul := map_natCast_smul diff --git a/Mathlib/Algebra/Module/Defs.lean b/Mathlib/Algebra/Module/Defs.lean index c51d03d323fe8..9b2be9d66ea8b 100644 --- a/Mathlib/Algebra/Module/Defs.lean +++ b/Mathlib/Algebra/Module/Defs.lean @@ -3,10 +3,6 @@ Copyright (c) 2015 Nathaniel Thomas. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Nathaniel Thomas, Jeremy Avigad, Johannes Hölzl, Mario Carneiro -/ -import Mathlib.Algebra.Group.Hom.End -import Mathlib.Algebra.GroupWithZero.Action.Units -import Mathlib.Algebra.Ring.Invertible -import Mathlib.Algebra.Ring.Opposite import Mathlib.Algebra.SMulWithZero import Mathlib.Data.Int.Cast.Lemmas @@ -37,10 +33,11 @@ to use a canonical `Module` typeclass throughout. semimodule, module, vector space -/ +assert_not_exists Field +assert_not_exists Invertible assert_not_exists Multiset -assert_not_exists Set.indicator assert_not_exists Pi.single_smul₀ -assert_not_exists Field +assert_not_exists Set.indicator open Function Set @@ -81,10 +78,6 @@ instance AddCommGroup.toNatModule : Module ℕ M where zero_smul := zero_nsmul add_smul r s x := add_nsmul x r s -theorem AddMonoid.End.natCast_def (n : ℕ) : - (↑n : AddMonoid.End M) = DistribMulAction.toAddMonoidEnd ℕ M n := - rfl - theorem add_smul : (r + s) • x = r • x + s • x := Module.add_smul r s x @@ -96,11 +89,6 @@ variable (R) -- Porting note: this is the letter of the mathlib3 version, but not really the spirit theorem two_smul : (2 : R) • x = x + x := by rw [← one_add_one_eq_two, add_smul, one_smul] -@[simp] -theorem invOf_two_smul_add_invOf_two_smul [Invertible (2 : R)] (x : M) : - (⅟ 2 : R) • x + (⅟ 2 : R) • x = x := - Convex.combo_self invOf_two_add_invOf_two _ - /-- Pullback a `Module` structure along an injective additive monoid homomorphism. See note [reducible non-instances]. -/ protected abbrev Function.Injective.module [AddCommMonoid M₂] [SMul R M₂] (f : M₂ →+ M) @@ -145,32 +133,7 @@ abbrev Module.compHom [Semiring S] (f : S →+* R) : Module S M := -- https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/Heterogeneous.20scalar.20multiplication add_smul := fun r s x => show f (r + s) • x = f r • x + f s • x by simp [add_smul] } -variable (R) - -/-- `(•)` as an `AddMonoidHom`. - -This is a stronger version of `DistribMulAction.toAddMonoidEnd` -/ -@[simps! apply_apply] -def Module.toAddMonoidEnd : R →+* AddMonoid.End M := - { DistribMulAction.toAddMonoidEnd R M with - -- Porting note: the two `show`s weren't needed in mathlib3. - -- Somehow, now that `SMul` is heterogeneous, it can't unfold earlier fields of a definition for - -- use in later fields. See - -- https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/Heterogeneous.20scalar.20multiplication - map_zero' := AddMonoidHom.ext fun r => show (0 : R) • r = 0 by simp - map_add' := fun x y => - AddMonoidHom.ext fun r => show (x + y) • r = x • r + y • r by simp [add_smul] } - -/-- A convenience alias for `Module.toAddMonoidEnd` as an `AddMonoidHom`, usually to allow the -use of `AddMonoidHom.flip`. -/ -def smulAddHom : R →+ M →+ M := - (Module.toAddMonoidEnd R M).toAddMonoidHom - -variable {R M} - -@[simp] -theorem smulAddHom_apply (r : R) (x : M) : smulAddHom R M r x = r • x := - rfl +variable {M} theorem Module.eq_zero_of_zero_eq_one (zero_eq_one : (0 : R) = 1) : x = 0 := by rw [← one_smul R x, ← zero_eq_one, zero_smul] @@ -194,10 +157,6 @@ instance AddCommGroup.toIntModule : Module ℤ M where zero_smul := zero_zsmul add_smul r s x := add_zsmul x r s -theorem AddMonoid.End.intCast_def (z : ℤ) : - (↑z : AddMonoid.End M) = DistribMulAction.toAddMonoidEnd ℤ M z := - rfl - variable {R M} theorem Convex.combo_eq_smul_sub_add [Module R M] {x y : M} {a b : R} (h : a + b = 1) : @@ -281,13 +240,6 @@ instance (priority := 910) Semiring.toModule [Semiring R] : Module R R where zero_smul := zero_mul smul_zero := mul_zero --- see Note [lower instance priority] -/-- Like `Semiring.toModule`, but multiplies on the right. -/ -instance (priority := 910) Semiring.toOppositeModule [Semiring R] : Module Rᵐᵒᵖ R := - { MonoidWithZero.toOppositeMulActionWithZero R with - smul_add := fun _ _ _ => add_mul _ _ _ - add_smul := fun _ _ _ => mul_add _ _ _ } - /-- A ring homomorphism `f : R →+* M` defines a module structure by `r • x = f r * x`. -/ def RingHom.toModule [Semiring R] [Semiring S] (f : R →+* S) : Module R S := Module.compHom S f @@ -358,56 +310,11 @@ instance AddCommMonoid.nat_isScalarTower : IsScalarTower ℕ R M where end AddCommMonoid -section AddCommGroup - -variable [Ring R] [AddCommGroup M] [Module R M] - -section - -variable (R) - -/-- `zsmul` is equal to any other module structure via a cast. -/ -lemma Int.cast_smul_eq_zsmul (n : ℤ) (b : M) : (n : R) • b = n • b := - have : ((smulAddHom R M).flip b).comp (Int.castAddHom R) = (smulAddHom ℤ M).flip b := by - apply AddMonoidHom.ext_int - simp - DFunLike.congr_fun this n - -@[deprecated (since := "2024-07-23")] alias intCast_smul := Int.cast_smul_eq_zsmul - -/-- `zsmul` is equal to any other module structure via a cast. -/ -@[deprecated Int.cast_smul_eq_zsmul (since := "2024-07-23")] -theorem zsmul_eq_smul_cast (n : ℤ) (b : M) : n • b = (n : R) • b := (Int.cast_smul_eq_zsmul ..).symm - -end - -/-- Convert back any exotic `ℤ`-smul to the canonical instance. This should not be needed since in -mathlib all `AddCommGroup`s should normally have exactly one `ℤ`-module structure by design. -/ -theorem int_smul_eq_zsmul (h : Module ℤ M) (n : ℤ) (x : M) : @SMul.smul ℤ M h.toSMul n x = n • x := - Int.cast_smul_eq_zsmul .. - -/-- All `ℤ`-module structures are equal. Not an instance since in mathlib all `AddCommGroup` -should normally have exactly one `ℤ`-module structure by design. -/ -def AddCommGroup.uniqueIntModule : Unique (Module ℤ M) where - default := by infer_instance - uniq P := (Module.ext' P _) fun n => by convert int_smul_eq_zsmul P n - -end AddCommGroup - -theorem map_intCast_smul [AddCommGroup M] [AddCommGroup M₂] {F : Type*} [FunLike F M M₂] - [AddMonoidHomClass F M M₂] (f : F) (R S : Type*) [Ring R] [Ring S] [Module R M] [Module S M₂] - (x : ℤ) (a : M) : - f ((x : R) • a) = (x : S) • f a := by simp only [Int.cast_smul_eq_zsmul, map_zsmul] - theorem map_natCast_smul [AddCommMonoid M] [AddCommMonoid M₂] {F : Type*} [FunLike F M M₂] [AddMonoidHomClass F M M₂] (f : F) (R S : Type*) [Semiring R] [Semiring S] [Module R M] [Module S M₂] (x : ℕ) (a : M) : f ((x : R) • a) = (x : S) • f a := by simp only [Nat.cast_smul_eq_nsmul, AddMonoidHom.map_nsmul, map_nsmul] -instance AddCommGroup.intIsScalarTower {R : Type u} {M : Type v} [Ring R] [AddCommGroup M] - [Module R M] : IsScalarTower ℤ R M where - smul_assoc n x y := ((smulAddHom R M).flip y).map_zsmul x n - -- Porting note (#10618): simp can prove this --@[simp] theorem Nat.smul_one_eq_cast {R : Type*} [NonAssocSemiring R] (m : ℕ) : m • (1 : R) = ↑m := by diff --git a/Mathlib/Algebra/Module/End.lean b/Mathlib/Algebra/Module/End.lean new file mode 100644 index 0000000000000..b1269c071ad3d --- /dev/null +++ b/Mathlib/Algebra/Module/End.lean @@ -0,0 +1,117 @@ +/- +Copyright (c) 2015 Nathaniel Thomas. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Nathaniel Thomas, Jeremy Avigad, Johannes Hölzl, Mario Carneiro +-/ +import Mathlib.Algebra.Group.Hom.End +import Mathlib.Algebra.Module.Defs + +/-! +# Module structure and endomorphisms + +In this file, we define `Module.toAddMonoidEnd`, which is `(•)` as a monoid homomorphism. +We use this to prove some results on scalar multiplication by integers. +-/ + +assert_not_exists Multiset +assert_not_exists Set.indicator +assert_not_exists Pi.single_smul₀ +assert_not_exists Field + +open Function Set + +universe u v + +variable {R S M M₂ : Type*} + +section AddCommMonoid + +variable [Semiring R] [AddCommMonoid M] [Module R M] (r s : R) (x : M) + +theorem AddMonoid.End.natCast_def (n : ℕ) : + (↑n : AddMonoid.End M) = DistribMulAction.toAddMonoidEnd ℕ M n := + rfl + +variable (R M) + +/-- `(•)` as an `AddMonoidHom`. + +This is a stronger version of `DistribMulAction.toAddMonoidEnd` -/ +@[simps! apply_apply] +def Module.toAddMonoidEnd : R →+* AddMonoid.End M := + { DistribMulAction.toAddMonoidEnd R M with + -- Porting note: the two `show`s weren't needed in mathlib3. + -- Somehow, now that `SMul` is heterogeneous, it can't unfold earlier fields of a definition for + -- use in later fields. See + -- https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/Heterogeneous.20scalar.20multiplication + map_zero' := AddMonoidHom.ext fun r => show (0 : R) • r = 0 by simp + map_add' := fun x y => + AddMonoidHom.ext fun r => show (x + y) • r = x • r + y • r by simp [add_smul] } + +/-- A convenience alias for `Module.toAddMonoidEnd` as an `AddMonoidHom`, usually to allow the +use of `AddMonoidHom.flip`. -/ +def smulAddHom : R →+ M →+ M := + (Module.toAddMonoidEnd R M).toAddMonoidHom + +variable {R M} + +@[simp] +theorem smulAddHom_apply (r : R) (x : M) : smulAddHom R M r x = r • x := + rfl + +end AddCommMonoid + +section AddCommGroup + +variable (R M) [Semiring R] [AddCommGroup M] + +theorem AddMonoid.End.intCast_def (z : ℤ) : + (↑z : AddMonoid.End M) = DistribMulAction.toAddMonoidEnd ℤ M z := + rfl + +end AddCommGroup + +section AddCommGroup + +variable [Ring R] [AddCommGroup M] [Module R M] + +section + +variable (R) + +/-- `zsmul` is equal to any other module structure via a cast. -/ +lemma Int.cast_smul_eq_zsmul (n : ℤ) (b : M) : (n : R) • b = n • b := + have : ((smulAddHom R M).flip b).comp (Int.castAddHom R) = (smulAddHom ℤ M).flip b := by + apply AddMonoidHom.ext_int + simp + DFunLike.congr_fun this n + +@[deprecated (since := "2024-07-23")] alias intCast_smul := Int.cast_smul_eq_zsmul + +/-- `zsmul` is equal to any other module structure via a cast. -/ +@[deprecated Int.cast_smul_eq_zsmul (since := "2024-07-23")] +theorem zsmul_eq_smul_cast (n : ℤ) (b : M) : n • b = (n : R) • b := (Int.cast_smul_eq_zsmul ..).symm + +end + +/-- Convert back any exotic `ℤ`-smul to the canonical instance. This should not be needed since in +mathlib all `AddCommGroup`s should normally have exactly one `ℤ`-module structure by design. -/ +theorem int_smul_eq_zsmul (h : Module ℤ M) (n : ℤ) (x : M) : @SMul.smul ℤ M h.toSMul n x = n • x := + Int.cast_smul_eq_zsmul .. + +/-- All `ℤ`-module structures are equal. Not an instance since in mathlib all `AddCommGroup` +should normally have exactly one `ℤ`-module structure by design. -/ +def AddCommGroup.uniqueIntModule : Unique (Module ℤ M) where + default := by infer_instance + uniq P := (Module.ext' P _) fun n => by convert int_smul_eq_zsmul P n + +end AddCommGroup + +theorem map_intCast_smul [AddCommGroup M] [AddCommGroup M₂] {F : Type*} [FunLike F M M₂] + [AddMonoidHomClass F M M₂] (f : F) (R S : Type*) [Ring R] [Ring S] [Module R M] [Module S M₂] + (x : ℤ) (a : M) : + f ((x : R) • a) = (x : S) • f a := by simp only [Int.cast_smul_eq_zsmul, map_zsmul] + +instance AddCommGroup.intIsScalarTower {R : Type u} {M : Type v} [Ring R] [AddCommGroup M] + [Module R M] : IsScalarTower ℤ R M where + smul_assoc n x y := ((smulAddHom R M).flip y).map_zsmul x n diff --git a/Mathlib/Algebra/Module/Equiv/Basic.lean b/Mathlib/Algebra/Module/Equiv/Basic.lean index 35af015470723..b5a356bcc0795 100644 --- a/Mathlib/Algebra/Module/Equiv/Basic.lean +++ b/Mathlib/Algebra/Module/Equiv/Basic.lean @@ -6,6 +6,7 @@ Authors: Nathaniel Thomas, Jeremy Avigad, Johannes Hölzl, Mario Carneiro, Anne -/ import Mathlib.Algebra.Field.Defs import Mathlib.Algebra.GroupWithZero.Action.Basic +import Mathlib.Algebra.GroupWithZero.Action.Units import Mathlib.Algebra.Module.Equiv.Defs import Mathlib.Algebra.Module.Hom import Mathlib.Algebra.Module.LinearMap.End diff --git a/Mathlib/Algebra/Module/Opposites.lean b/Mathlib/Algebra/Module/Equiv/Opposite.lean similarity index 73% rename from Mathlib/Algebra/Module/Opposites.lean rename to Mathlib/Algebra/Module/Equiv/Opposite.lean index 914d9f3145572..1412025bff29f 100644 --- a/Mathlib/Algebra/Module/Opposites.lean +++ b/Mathlib/Algebra/Module/Equiv/Opposite.lean @@ -3,8 +3,8 @@ Copyright (c) 2020 Eric Wieser. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Eric Wieser -/ -import Mathlib.Algebra.GroupWithZero.Action.Opposite import Mathlib.Algebra.Module.Equiv.Defs +import Mathlib.Algebra.Module.Opposite /-! # Module operations on `Mᵐᵒᵖ` @@ -13,6 +13,23 @@ This file contains definitions that build on top of the group action definitions `Mathlib.Algebra.GroupWithZero.Action.Opposite`. -/ +section + +variable {R S M : Type*} [Semiring R] [Semiring S] [AddCommMonoid M] [Module S M] + +@[ext high] +theorem LinearMap.ext_ring_op + {σ : Rᵐᵒᵖ →+* S} {f g : R →ₛₗ[σ] M} (h : f (1 : R) = g (1 : R)) : + f = g := + ext fun x ↦ by + -- Porting note: replaced the oneliner `rw` proof with a partially term-mode proof + -- because `rw` was giving "motive is type incorrect" errors + rw [← one_mul x, ← op_smul_eq_mul] + refine (f.map_smulₛₗ (MulOpposite.op x) 1).trans ?_ + rw [h] + exact (g.map_smulₛₗ (MulOpposite.op x) 1).symm + +end namespace MulOpposite @@ -20,11 +37,6 @@ universe u v variable (R : Type u) {M : Type v} [Semiring R] [AddCommMonoid M] [Module R M] -/-- `MulOpposite.distribMulAction` extends to a `Module` -/ -instance instModule : Module R Mᵐᵒᵖ where - add_smul _ _ _ := unop_injective <| add_smul _ _ _ - zero_smul _ := unop_injective <| zero_smul _ _ - /-- The function `op` is a linear equivalence. -/ def opLinearEquiv : M ≃ₗ[R] Mᵐᵒᵖ := { opAddEquiv with map_smul' := MulOpposite.op_smul } diff --git a/Mathlib/Algebra/Module/Hom.lean b/Mathlib/Algebra/Module/Hom.lean index 19bee131fe2dd..66bdfade35658 100644 --- a/Mathlib/Algebra/Module/Hom.lean +++ b/Mathlib/Algebra/Module/Hom.lean @@ -3,8 +3,9 @@ Copyright (c) 2021 Eric Wieser. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Eric Wieser -/ -import Mathlib.Algebra.Module.Defs import Mathlib.Algebra.Group.Hom.Instances +import Mathlib.Algebra.Module.End +import Mathlib.Algebra.Ring.Opposite import Mathlib.GroupTheory.GroupAction.DomAct.Basic /-! diff --git a/Mathlib/Algebra/Module/LinearMap/Basic.lean b/Mathlib/Algebra/Module/LinearMap/Basic.lean index baff9dbf891d4..b82258f2ecc03 100644 --- a/Mathlib/Algebra/Module/LinearMap/Basic.lean +++ b/Mathlib/Algebra/Module/LinearMap/Basic.lean @@ -6,6 +6,7 @@ Authors: Nathaniel Thomas, Jeremy Avigad, Johannes Hölzl, Mario Carneiro, Anne -/ import Mathlib.Algebra.Module.LinearMap.Defs import Mathlib.Algebra.NoZeroSMulDivisors.Pi +import Mathlib.Algebra.Ring.Opposite import Mathlib.GroupTheory.GroupAction.DomAct.Basic /-! diff --git a/Mathlib/Algebra/Module/LinearMap/Defs.lean b/Mathlib/Algebra/Module/LinearMap/Defs.lean index 21a6e98c5445a..8e3800a3eed03 100644 --- a/Mathlib/Algebra/Module/LinearMap/Defs.lean +++ b/Mathlib/Algebra/Module/LinearMap/Defs.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Nathaniel Thomas, Jeremy Avigad, Johannes Hölzl, Mario Carneiro, Anne Baanen, Frédéric Dupuis, Heather Macbeth -/ -import Mathlib.Algebra.Module.Defs +import Mathlib.Algebra.Group.Hom.Instances import Mathlib.Algebra.Ring.CompTypeclasses import Mathlib.GroupTheory.GroupAction.Hom @@ -448,17 +448,6 @@ theorem toAddMonoidHom_injective : theorem ext_ring {f g : R →ₛₗ[σ] M₃} (h : f 1 = g 1) : f = g := ext fun x ↦ by rw [← mul_one x, ← smul_eq_mul, f.map_smulₛₗ, g.map_smulₛₗ, h] -@[ext high] -theorem ext_ring_op {σ : Rᵐᵒᵖ →+* S} {f g : R →ₛₗ[σ] M₃} (h : f (1 : R) = g (1 : R)) : - f = g := - ext fun x ↦ by - -- Porting note: replaced the oneliner `rw` proof with a partially term-mode proof - -- because `rw` was giving "motive is type incorrect" errors - rw [← one_mul x, ← op_smul_eq_mul] - refine (f.map_smulₛₗ (MulOpposite.op x) 1).trans ?_ - rw [h] - exact (g.map_smulₛₗ (MulOpposite.op x) 1).symm - end /-- Interpret a `RingHom` `f` as an `f`-semilinear map. -/ diff --git a/Mathlib/Algebra/Module/LinearMap/End.lean b/Mathlib/Algebra/Module/LinearMap/End.lean index a861127390ad8..f2c1f9a46d3d6 100644 --- a/Mathlib/Algebra/Module/LinearMap/End.lean +++ b/Mathlib/Algebra/Module/LinearMap/End.lean @@ -6,6 +6,7 @@ Authors: Nathaniel Thomas, Jeremy Avigad, Johannes Hölzl, Mario Carneiro, Anne -/ import Mathlib.Algebra.GroupPower.IterateHom import Mathlib.Algebra.Module.LinearMap.Defs +import Mathlib.Algebra.Module.Equiv.Opposite import Mathlib.Algebra.NoZeroSMulDivisors.Defs /-! diff --git a/Mathlib/Algebra/Module/Opposite.lean b/Mathlib/Algebra/Module/Opposite.lean new file mode 100644 index 0000000000000..5994afcd97b8f --- /dev/null +++ b/Mathlib/Algebra/Module/Opposite.lean @@ -0,0 +1,43 @@ +/- +Copyright (c) 2020 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser +-/ +import Mathlib.Algebra.GroupWithZero.Action.Opposite +import Mathlib.Algebra.Module.Defs +import Mathlib.Algebra.Ring.Opposite + +/-! +# Module operations on `Mᵐᵒᵖ` + +This file contains definitions that build on top of the group action definitions in +`Mathlib.Algebra.GroupWithZero.Action.Opposite`. +-/ + +assert_not_exists LinearMap + +section + +variable {R S M : Type*} [Semiring R] [Semiring S] [AddCommMonoid M] [Module S M] + +-- see Note [lower instance priority] +/-- Like `Semiring.toModule`, but multiplies on the right. -/ +instance (priority := 910) Semiring.toOppositeModule [Semiring R] : Module Rᵐᵒᵖ R := + { MonoidWithZero.toOppositeMulActionWithZero R with + smul_add := fun _ _ _ => add_mul _ _ _ + add_smul := fun _ _ _ => mul_add _ _ _ } + +end + +namespace MulOpposite + +universe u v + +variable (R : Type u) {M : Type v} [Semiring R] [AddCommMonoid M] [Module R M] + +/-- `MulOpposite.distribMulAction` extends to a `Module` -/ +instance instModule : Module R Mᵐᵒᵖ where + add_smul _ _ _ := unop_injective <| add_smul _ _ _ + zero_smul _ := unop_injective <| zero_smul _ _ + +end MulOpposite diff --git a/Mathlib/Algebra/NoZeroSMulDivisors/Basic.lean b/Mathlib/Algebra/NoZeroSMulDivisors/Basic.lean index f33114e9583a3..c4b95955f40ac 100644 --- a/Mathlib/Algebra/NoZeroSMulDivisors/Basic.lean +++ b/Mathlib/Algebra/NoZeroSMulDivisors/Basic.lean @@ -3,7 +3,8 @@ Copyright (c) 2015 Nathaniel Thomas. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Anne Baanen, Yury Kudryashov, Joseph Myers, Heather Macbeth, Kim Morrison, Yaël Dillies -/ -import Mathlib.Algebra.Module.Defs +import Mathlib.Algebra.GroupWithZero.Action.Units +import Mathlib.Algebra.Module.End import Mathlib.Algebra.NoZeroSMulDivisors.Defs /-! diff --git a/Mathlib/Algebra/Periodic.lean b/Mathlib/Algebra/Periodic.lean index 995e7e734029d..1225d22fde07c 100644 --- a/Mathlib/Algebra/Periodic.lean +++ b/Mathlib/Algebra/Periodic.lean @@ -6,8 +6,9 @@ Authors: Benjamin Davidson import Mathlib.Algebra.Field.Opposite import Mathlib.Algebra.Group.Subgroup.ZPowers import Mathlib.Algebra.Group.Submonoid.Membership -import Mathlib.Algebra.Ring.NegOnePow +import Mathlib.Algebra.Module.Opposite import Mathlib.Algebra.Order.Archimedean.Basic +import Mathlib.Algebra.Ring.NegOnePow import Mathlib.GroupTheory.Coset.Card /-! diff --git a/Mathlib/Algebra/Ring/AddAut.lean b/Mathlib/Algebra/Ring/AddAut.lean index f593f12d75bda..8a75d5e9bcde8 100644 --- a/Mathlib/Algebra/Ring/AddAut.lean +++ b/Mathlib/Algebra/Ring/AddAut.lean @@ -4,7 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import Mathlib.Algebra.GroupWithZero.Action.Basic -import Mathlib.Algebra.Module.Defs +import Mathlib.Algebra.GroupWithZero.Action.Units +import Mathlib.Algebra.Module.Opposite /-! # Multiplication on the left/right as additive automorphisms diff --git a/Mathlib/Algebra/Ring/Subsemiring/MulOpposite.lean b/Mathlib/Algebra/Ring/Subsemiring/MulOpposite.lean index f709792084761..2b5aa3ce3ce5e 100644 --- a/Mathlib/Algebra/Ring/Subsemiring/MulOpposite.lean +++ b/Mathlib/Algebra/Ring/Subsemiring/MulOpposite.lean @@ -5,6 +5,7 @@ Authors: Jz Pan -/ import Mathlib.Algebra.Group.Submonoid.MulOpposite import Mathlib.Algebra.Ring.Subsemiring.Basic +import Mathlib.Algebra.Ring.Opposite /-! diff --git a/Mathlib/CategoryTheory/Linear/Basic.lean b/Mathlib/CategoryTheory/Linear/Basic.lean index 5f724dcf36cb9..b4d90d75d19a2 100644 --- a/Mathlib/CategoryTheory/Linear/Basic.lean +++ b/Mathlib/CategoryTheory/Linear/Basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison -/ import Mathlib.Algebra.Algebra.Defs +import Mathlib.Algebra.Group.Invertible.Defs import Mathlib.Algebra.Module.Equiv.Defs import Mathlib.CategoryTheory.Preadditive.Basic diff --git a/Mathlib/CategoryTheory/Preadditive/Basic.lean b/Mathlib/CategoryTheory/Preadditive/Basic.lean index 8df1e560ccfc4..0ddcb0081738b 100644 --- a/Mathlib/CategoryTheory/Preadditive/Basic.lean +++ b/Mathlib/CategoryTheory/Preadditive/Basic.lean @@ -5,7 +5,7 @@ Authors: Markus Himmel, Jakob von Raumer -/ import Mathlib.Algebra.BigOperators.Group.Finset import Mathlib.Algebra.Group.Hom.Defs -import Mathlib.Algebra.Module.Defs +import Mathlib.Algebra.Module.End import Mathlib.CategoryTheory.Endomorphism import Mathlib.CategoryTheory.Limits.Shapes.Kernels diff --git a/Mathlib/Data/Set/Pointwise/SMul.lean b/Mathlib/Data/Set/Pointwise/SMul.lean index bd562b3fb02ad..68200c4096290 100644 --- a/Mathlib/Data/Set/Pointwise/SMul.lean +++ b/Mathlib/Data/Set/Pointwise/SMul.lean @@ -7,6 +7,7 @@ import Mathlib.Algebra.Group.Pi.Basic import Mathlib.Algebra.Group.Pointwise.Set.Basic import Mathlib.Algebra.GroupWithZero.Action.Basic import Mathlib.Algebra.Module.Defs +import Mathlib.Algebra.Ring.Opposite import Mathlib.Algebra.NoZeroSMulDivisors.Defs import Mathlib.Data.Set.Pairwise.Basic diff --git a/Mathlib/Data/ZMod/Basic.lean b/Mathlib/Data/ZMod/Basic.lean index cfb7aa79acee7..4c03f03635c6f 100644 --- a/Mathlib/Data/ZMod/Basic.lean +++ b/Mathlib/Data/ZMod/Basic.lean @@ -3,6 +3,7 @@ Copyright (c) 2018 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes -/ +import Mathlib.Algebra.Module.End import Mathlib.Algebra.Ring.Prod import Mathlib.GroupTheory.GroupAction.SubMulAction import Mathlib.GroupTheory.OrderOfElement diff --git a/Mathlib/LinearAlgebra/AffineSpace/Midpoint.lean b/Mathlib/LinearAlgebra/AffineSpace/Midpoint.lean index d57b4f3a3c569..90be3fd7ffbba 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/Midpoint.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/Midpoint.lean @@ -3,6 +3,7 @@ Copyright (c) 2020 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ +import Mathlib.Algebra.Module.Basic import Mathlib.LinearAlgebra.AffineSpace.AffineEquiv /-! diff --git a/Mathlib/LinearAlgebra/CliffordAlgebra/Conjugation.lean b/Mathlib/LinearAlgebra/CliffordAlgebra/Conjugation.lean index 019e24cf0951f..52780073fdfe0 100644 --- a/Mathlib/LinearAlgebra/CliffordAlgebra/Conjugation.lean +++ b/Mathlib/LinearAlgebra/CliffordAlgebra/Conjugation.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Eric Wieser -/ import Mathlib.LinearAlgebra.CliffordAlgebra.Grading -import Mathlib.Algebra.Module.Opposites +import Mathlib.Algebra.Module.Opposite /-! # Conjugations diff --git a/Mathlib/RingTheory/OreLocalization/Ring.lean b/Mathlib/RingTheory/OreLocalization/Ring.lean index 73dca39178779..d76032bfbc512 100644 --- a/Mathlib/RingTheory/OreLocalization/Ring.lean +++ b/Mathlib/RingTheory/OreLocalization/Ring.lean @@ -4,9 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jakob von Raumer, Kevin Klinge, Andrew Yang -/ -import Mathlib.Algebra.GroupWithZero.NonZeroDivisors import Mathlib.Algebra.Algebra.Defs import Mathlib.Algebra.Field.Defs +import Mathlib.Algebra.GroupWithZero.NonZeroDivisors +import Mathlib.Algebra.Module.End import Mathlib.RingTheory.OreLocalization.Basic /-! diff --git a/Mathlib/RingTheory/TwoSidedIdeal/Basic.lean b/Mathlib/RingTheory/TwoSidedIdeal/Basic.lean index dd7cee192929b..532582ec88dd8 100644 --- a/Mathlib/RingTheory/TwoSidedIdeal/Basic.lean +++ b/Mathlib/RingTheory/TwoSidedIdeal/Basic.lean @@ -8,6 +8,7 @@ import Mathlib.Tactic.Abel import Mathlib.GroupTheory.GroupAction.SubMulAction import Mathlib.RingTheory.Congruence.Basic import Mathlib.Algebra.Module.LinearMap.Defs +import Mathlib.Algebra.Module.Opposite /-! # Two Sided Ideals diff --git a/Mathlib/Topology/Algebra/Module/Basic.lean b/Mathlib/Topology/Algebra/Module/Basic.lean index c2feb47fccf1a..6293e9dea3c44 100644 --- a/Mathlib/Topology/Algebra/Module/Basic.lean +++ b/Mathlib/Topology/Algebra/Module/Basic.lean @@ -4,15 +4,15 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jan-David Salchow, Sébastien Gouëzel, Jean Lo, Yury Kudryashov, Frédéric Dupuis, Heather Macbeth -/ -import Mathlib.Topology.Algebra.Ring.Basic +import Mathlib.Algebra.Algebra.Defs +import Mathlib.Algebra.Module.Opposite +import Mathlib.LinearAlgebra.Finsupp +import Mathlib.LinearAlgebra.Pi +import Mathlib.LinearAlgebra.Projection import Mathlib.Topology.Algebra.MulAction +import Mathlib.Topology.Algebra.Ring.Basic import Mathlib.Topology.Algebra.UniformGroup import Mathlib.Topology.UniformSpace.UniformEmbedding -import Mathlib.Algebra.Algebra.Defs -import Mathlib.LinearAlgebra.Projection -import Mathlib.LinearAlgebra.Pi -import Mathlib.LinearAlgebra.Finsupp -import Mathlib.Algebra.Module.Opposites /-! # Theory of topological modules and continuous linear maps. diff --git a/Mathlib/Topology/Algebra/UniformMulAction.lean b/Mathlib/Topology/Algebra/UniformMulAction.lean index 3f0de8b6107d6..7fb5a47c14347 100644 --- a/Mathlib/Topology/Algebra/UniformMulAction.lean +++ b/Mathlib/Topology/Algebra/UniformMulAction.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ +import Mathlib.Algebra.Module.Opposite import Mathlib.Topology.Algebra.UniformGroup import Mathlib.Topology.UniformSpace.Completion diff --git a/scripts/noshake.json b/scripts/noshake.json index 9cf7d5e4a2312..e6f8ba397178e 100644 --- a/scripts/noshake.json +++ b/scripts/noshake.json @@ -378,6 +378,7 @@ ["Mathlib.Algebra.Order.Field.Basic"], "Mathlib.CategoryTheory.Sites.IsSheafFor": ["Mathlib.CategoryTheory.Limits.Shapes.Pullback.Mono"], + "Mathlib.CategoryTheory.Preadditive.Basic": ["Mathlib.Algebra.Module.End"], "Mathlib.CategoryTheory.Monoidal.Rigid.Basic": ["Mathlib.Tactic.CategoryTheory.Monoidal.Basic"], "Mathlib.CategoryTheory.Monoidal.Braided.Basic": From a5a1a0dcb784f2801ebf84f26072b8f550194335 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Thu, 24 Oct 2024 15:52:23 +0000 Subject: [PATCH 376/505] feat(CategoryTheory): the short exact sequence of abelian sheaves of a Mayer-Vietoris square (#15053) The long exact sequence of Ext associated to this short exact sequence will give the Mayer-Vietoris long exact sequence in sheaf cohomology. --- Mathlib.lean | 1 + Mathlib/Algebra/Homology/Square.lean | 72 +++++++++++++++++++ .../CategoryTheory/Adjunction/Evaluation.lean | 8 +-- .../Sites/MayerVietorisSquare.lean | 65 +++++++++++++++-- 4 files changed, 137 insertions(+), 9 deletions(-) create mode 100644 Mathlib/Algebra/Homology/Square.lean diff --git a/Mathlib.lean b/Mathlib.lean index 2d79a052f5058..f74556466750a 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -422,6 +422,7 @@ import Mathlib.Algebra.Homology.ShortComplex.ShortExact import Mathlib.Algebra.Homology.ShortComplex.SnakeLemma import Mathlib.Algebra.Homology.Single import Mathlib.Algebra.Homology.SingleHomology +import Mathlib.Algebra.Homology.Square import Mathlib.Algebra.Homology.TotalComplex import Mathlib.Algebra.Homology.TotalComplexShift import Mathlib.Algebra.Homology.TotalComplexSymmetry diff --git a/Mathlib/Algebra/Homology/Square.lean b/Mathlib/Algebra/Homology/Square.lean new file mode 100644 index 0000000000000..d2c3faea2d9aa --- /dev/null +++ b/Mathlib/Algebra/Homology/Square.lean @@ -0,0 +1,72 @@ +/- +Copyright (c) 2024 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ +import Mathlib.Algebra.Homology.CommSq +import Mathlib.CategoryTheory.Limits.Shapes.Pullback.Square + +/-! +# Relation between pullback/pushout squares and kernel/cokernel sequences + +This file is the bundled counterpart of `Algebra.Homology.CommSq`. +The same results are obtained here for squares `sq : Square C` where +`C` is an additive category. + +-/ +namespace CategoryTheory + +open Category Limits + +namespace Square + +variable {C : Type*} [Category C] [Preadditive C] + (sq : Square C) [HasBinaryBiproduct sq.X₂ sq.X₃] + +/-- The cokernel cofork attached to a commutative square in a preadditive category. -/ +noncomputable abbrev cokernelCofork : + CokernelCofork (biprod.lift sq.f₁₂ (-sq.f₁₃)) := + CokernelCofork.ofπ (biprod.desc sq.f₂₄ sq.f₃₄) (by simp [sq.fac]) + +/-- A commutative square in a preadditive category is a pushout square iff +the corresponding diagram `X₁ ⟶ X₂ ⊞ X₃ ⟶ X₄ ⟶ 0` makes `X₄` a cokernel. -/ +noncomputable def isPushoutEquivIsColimitCokernelCofork : + sq.IsPushout ≃ IsColimit sq.cokernelCofork := + Equiv.trans + { toFun := fun h ↦ h.isColimit + invFun := fun h ↦ IsPushout.mk _ h + left_inv := fun _ ↦ rfl + right_inv := fun _ ↦ Subsingleton.elim _ _ } + sq.commSq.isColimitEquivIsColimitCokernelCofork + +variable {sq} in +/-- The colimit cokernel cofork attached to a pushout square. -/ +noncomputable def IsPushout.isColimitCokernelCofork (h : sq.IsPushout) : + IsColimit sq.cokernelCofork := + h.isColimitEquivIsColimitCokernelCofork h.isColimit + +/-- The kernel fork attached to a commutative square in a preadditive category. -/ +noncomputable abbrev kernelFork : + KernelFork (biprod.desc sq.f₂₄ (-sq.f₃₄)) := + KernelFork.ofι (biprod.lift sq.f₁₂ sq.f₁₃) (by simp [sq.fac]) + +/-- A commutative square in a preadditive category is a pullback square iff +the corresponding diagram `0 ⟶ X₁ ⟶ X₂ ⊞ X₃ ⟶ X₄ ⟶ 0` makes `X₁` a kernel. -/ +noncomputable def isPullbackEquivIsLimitKernelFork : + sq.IsPullback ≃ IsLimit sq.kernelFork := + Equiv.trans + { toFun := fun h ↦ h.isLimit + invFun := fun h ↦ IsPullback.mk _ h + left_inv := fun _ ↦ rfl + right_inv := fun _ ↦ Subsingleton.elim _ _ } + sq.commSq.isLimitEquivIsLimitKernelFork + +variable {sq} in +/-- The limit kernel fork attached to a pullback square. -/ +noncomputable def IsPullback.isLimitKernelFork (h : sq.IsPullback) : + IsLimit sq.kernelFork := + h.isLimitEquivIsLimitKernelFork h.isLimit + +end Square + +end CategoryTheory diff --git a/Mathlib/CategoryTheory/Adjunction/Evaluation.lean b/Mathlib/CategoryTheory/Adjunction/Evaluation.lean index d7491bf0972fe..2336e9442f689 100644 --- a/Mathlib/CategoryTheory/Adjunction/Evaluation.lean +++ b/Mathlib/CategoryTheory/Adjunction/Evaluation.lean @@ -19,9 +19,10 @@ namespace CategoryTheory open CategoryTheory.Limits -universe v₁ v₂ u₁ u₂ +universe v₁ v₂ v₃ u₁ u₂ u₃ variable {C : Type u₁} [Category.{v₁} C] (D : Type u₂) [Category.{v₂} D] + (E : Type u₃) [Category.{v₃} E] noncomputable section @@ -73,7 +74,7 @@ def evaluationAdjunctionRight (c : C) : evaluationLeftAdjoint D c ⊣ (evaluatio instance evaluationIsRightAdjoint (c : C) : ((evaluation _ D).obj c).IsRightAdjoint := ⟨_, ⟨evaluationAdjunctionRight _ _⟩⟩ -/-- See also the file `CategoryTheory.Limits.FunctorCategoryEpiMono` +/-- See also the file `CategoryTheory.Limits.FunctorCategory.EpiMono` for a similar result under a `HasPullbacks` assumption. -/ theorem NatTrans.mono_iff_mono_app' {F G : C ⥤ D} (η : F ⟶ G) : Mono η ↔ ∀ c, Mono (η.app c) := by constructor @@ -130,8 +131,7 @@ def evaluationAdjunctionLeft (c : C) : (evaluation _ _).obj c ⊣ evaluationRigh instance evaluationIsLeftAdjoint (c : C) : ((evaluation _ D).obj c).IsLeftAdjoint := ⟨_, ⟨evaluationAdjunctionLeft _ _⟩⟩ - -/-- See also the file `CategoryTheory.Limits.FunctorCategoryEpiMono` +/-- See also the file `CategoryTheory.Limits.FunctorCategory.EpiMono` for a similar result under a `HasPushouts` assumption. -/ theorem NatTrans.epi_iff_epi_app' {F G : C ⥤ D} (η : F ⟶ G) : Epi η ↔ ∀ c, Epi (η.app c) := by constructor diff --git a/Mathlib/CategoryTheory/Sites/MayerVietorisSquare.lean b/Mathlib/CategoryTheory/Sites/MayerVietorisSquare.lean index c36e6664261b4..2d994cb919f36 100644 --- a/Mathlib/CategoryTheory/Sites/MayerVietorisSquare.lean +++ b/Mathlib/CategoryTheory/Sites/MayerVietorisSquare.lean @@ -3,10 +3,15 @@ Copyright (c) 2024 Joël Riou. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Joël Riou -/ +import Mathlib.Algebra.Category.Grp.Abelian import Mathlib.Algebra.Category.Grp.Adjunctions -import Mathlib.CategoryTheory.Sites.Adjunction +import Mathlib.Algebra.Homology.ShortComplex.ShortExact +import Mathlib.Algebra.Homology.Square +import Mathlib.CategoryTheory.Limits.FunctorCategory.EpiMono import Mathlib.CategoryTheory.Limits.Preserves.Shapes.Square import Mathlib.CategoryTheory.Limits.Shapes.Types +import Mathlib.CategoryTheory.Sites.Abelian +import Mathlib.CategoryTheory.Sites.Adjunction import Mathlib.CategoryTheory.Sites.Sheafification /-! @@ -52,10 +57,10 @@ namespace CategoryTheory open Limits Opposite variable {C : Type u} [Category.{v} C] - {J : GrothendieckTopology C} [HasWeakSheafify J (Type v)] + {J : GrothendieckTopology C} -@[simp] lemma Sheaf.isPullback_square_op_map_yoneda_presheafToSheaf_yoneda_iff + [HasWeakSheafify J (Type v)] (F : Sheaf J (Type v)) (sq : Square C) : (sq.op.map ((yoneda ⋙ presheafToSheaf J _).op ⋙ yoneda.obj F)).IsPullback ↔ (sq.op.map F.val).IsPullback := by @@ -79,17 +84,23 @@ variable (J) topology consists of a commutative square `f₁₂ ≫ f₂₄ = f₁₃ ≫ f₃₄` in `C` such that `f₁₃` is a monomorphism and that the square becomes a pushout square in the category of sheaves of sets. -/ -structure MayerVietorisSquare extends Square C where +structure MayerVietorisSquare [HasWeakSheafify J (Type v)] extends Square C where mono_f₁₃ : Mono toSquare.f₁₃ := by infer_instance /-- the square becomes a pushout square in the category of sheaves of types -/ isPushout : (toSquare.map (yoneda ⋙ presheafToSheaf J _)).IsPushout namespace MayerVietorisSquare +attribute [instance] mono_f₁₃ + variable {J} +section + +variable [HasWeakSheafify J (Type v)] + /-- Constructor for Mayer-Vietoris squares taking as an input -a square `sq` such that `sq.f₂₄` is a mono and that for every +a square `sq` such that `sq.f₁₃` is a mono and that for every sheaf of types `F`, the square `sq.op.map F.val` is a pullback square. -/ @[simps toSquare] noncomputable def mk' (sq : Square C) [Mono sq.f₁₃] @@ -206,6 +217,50 @@ lemma sheafCondition_of_sheaf {A : Type u'} [Category.{v} A] (S.isPushout.op.map (yoneda.obj ⟨_, (isSheaf_iff_isSheaf_of_type _ _).2 (F.cond X.unop)⟩)) +end + +variable [HasWeakSheafify J (Type v)] [HasSheafify J AddCommGrp.{v}] + (S : J.MayerVietorisSquare) + +/-- The short complex of abelian sheaves +`ℤ[S.X₁] ⟶ ℤ[S.X₂] ⊞ ℤ[S.X₃] ⟶ ℤ[S.X₄]` +where the left map is a difference and the right map a sum. -/ +@[simps] +noncomputable def shortComplex : + ShortComplex (Sheaf J AddCommGrp.{v}) where + X₁ := (presheafToSheaf J _).obj (yoneda.obj S.X₁ ⋙ AddCommGrp.free) + X₂ := (presheafToSheaf J _).obj (yoneda.obj S.X₂ ⋙ AddCommGrp.free) ⊞ + (presheafToSheaf J _).obj (yoneda.obj S.X₃ ⋙ AddCommGrp.free) + X₃ := (presheafToSheaf J _).obj (yoneda.obj S.X₄ ⋙ AddCommGrp.free) + f := + biprod.lift + ((presheafToSheaf J _).map (whiskerRight (yoneda.map S.f₁₂) _)) + (-(presheafToSheaf J _).map (whiskerRight (yoneda.map S.f₁₃) _)) + g := + biprod.desc + ((presheafToSheaf J _).map (whiskerRight (yoneda.map S.f₂₄) _)) + ((presheafToSheaf J _).map (whiskerRight (yoneda.map S.f₃₄) _)) + zero := (S.map (yoneda ⋙ (whiskeringRight _ _ _).obj AddCommGrp.free ⋙ + presheafToSheaf J _)).cokernelCofork.condition + +instance : Mono S.shortComplex.f := by + have : Mono (S.shortComplex.f ≫ biprod.snd) := by + dsimp + simp only [biprod.lift_snd] + infer_instance + exact mono_of_mono _ biprod.snd + +instance : Epi S.shortComplex.g := + (S.shortComplex.exact_and_epi_g_iff_g_is_cokernel.2 + ⟨S.isPushoutAddCommGrpFreeSheaf.isColimitCokernelCofork⟩).2 + +lemma shortComplex_exact : S.shortComplex.Exact := + ShortComplex.exact_of_g_is_cokernel _ + S.isPushoutAddCommGrpFreeSheaf.isColimitCokernelCofork + +lemma shortComplex_shortExact : S.shortComplex.ShortExact where + exact := S.shortComplex_exact + end MayerVietorisSquare end GrothendieckTopology From e1f8635a93689e15cc8cd292ace6de89bf5036f8 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Thu, 24 Oct 2024 15:52:24 +0000 Subject: [PATCH 377/505] refactor: create `CStarAlgebra` classes and refactor to use them (#16953) --- Mathlib.lean | 4 + Mathlib/Analysis/CStarAlgebra/Classes.lean | 104 ++++++++++++++++++ .../ContinuousFunctionalCalculus/Basic.lean | 8 +- .../Instances.lean | 31 ++---- .../ContinuousFunctionalCalculus/Order.lean | 11 +- .../CStarAlgebra/ContinuousLinearMap.lean | 17 +++ .../Analysis/CStarAlgebra/ContinuousMap.lean | 61 ++++++++++ .../Analysis/CStarAlgebra/GelfandDuality.lean | 8 +- Mathlib/Analysis/CStarAlgebra/Hom.lean | 12 +- .../CStarAlgebra/Module/Constructions.lean | 12 +- .../Analysis/CStarAlgebra/Module/Defs.lean | 29 +++-- Mathlib/Analysis/CStarAlgebra/Multiplier.lean | 3 + Mathlib/Analysis/CStarAlgebra/Spectrum.lean | 18 +-- .../Analysis/CStarAlgebra/Unitization.lean | 22 +++- Mathlib/Analysis/CStarAlgebra/lpSpace.lean | 35 ++++++ .../Analysis/InnerProductSpace/Adjoint.lean | 2 + .../Analysis/InnerProductSpace/StarOrder.lean | 1 + Mathlib/Analysis/Normed/Lp/lpSpace.lean | 12 +- .../ContinuousFunctionalCalculus/ExpLog.lean | 3 +- Mathlib/Analysis/VonNeumannAlgebra/Basic.lean | 6 +- Mathlib/Topology/ContinuousMap/Compact.lean | 4 +- 21 files changed, 299 insertions(+), 104 deletions(-) create mode 100644 Mathlib/Analysis/CStarAlgebra/Classes.lean create mode 100644 Mathlib/Analysis/CStarAlgebra/ContinuousLinearMap.lean create mode 100644 Mathlib/Analysis/CStarAlgebra/ContinuousMap.lean create mode 100644 Mathlib/Analysis/CStarAlgebra/lpSpace.lean diff --git a/Mathlib.lean b/Mathlib.lean index f74556466750a..34a16720097ff 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -990,6 +990,7 @@ import Mathlib.Analysis.BoxIntegral.Partition.Split import Mathlib.Analysis.BoxIntegral.Partition.SubboxInduction import Mathlib.Analysis.BoxIntegral.Partition.Tagged import Mathlib.Analysis.CStarAlgebra.Basic +import Mathlib.Analysis.CStarAlgebra.Classes import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Basic import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Instances import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Integral @@ -999,6 +1000,8 @@ import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Restrict import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Unique import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Unital import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Unitary +import Mathlib.Analysis.CStarAlgebra.ContinuousLinearMap +import Mathlib.Analysis.CStarAlgebra.ContinuousMap import Mathlib.Analysis.CStarAlgebra.Exponential import Mathlib.Analysis.CStarAlgebra.GelfandDuality import Mathlib.Analysis.CStarAlgebra.Hom @@ -1009,6 +1012,7 @@ import Mathlib.Analysis.CStarAlgebra.Module.Synonym import Mathlib.Analysis.CStarAlgebra.Multiplier import Mathlib.Analysis.CStarAlgebra.Spectrum import Mathlib.Analysis.CStarAlgebra.Unitization +import Mathlib.Analysis.CStarAlgebra.lpSpace import Mathlib.Analysis.Calculus.AddTorsor.AffineMap import Mathlib.Analysis.Calculus.AddTorsor.Coord import Mathlib.Analysis.Calculus.BumpFunction.Basic diff --git a/Mathlib/Analysis/CStarAlgebra/Classes.lean b/Mathlib/Analysis/CStarAlgebra/Classes.lean new file mode 100644 index 0000000000000..274d602111ef5 --- /dev/null +++ b/Mathlib/Analysis/CStarAlgebra/Classes.lean @@ -0,0 +1,104 @@ +/- +Copyright (c) 2024 Jireh Loreaux. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jireh Loreaux +-/ + +import Mathlib.Analysis.Complex.Basic +import Mathlib.Algebra.Star.NonUnitalSubalgebra + +/-! # Classes of C⋆-algebras + +This file defines classes for complex C⋆-algebras. These are (unital or non-unital, commutative or +noncommutative) Banach algebra over `ℂ` with an antimultiplicative conjugate-linear involution +(`star`) satisfying the C⋆-identity `∥star x * x∥ = ∥x∥ ^ 2`. + +## Notes + +These classes are not defined in `Mathlib.Analysis.CStarAlgebra.Basic` because they require +heavier imports. + +-/ + +/-- The class of non-unital (complex) C⋆-algebras. -/ +class NonUnitalCStarAlgebra (A : Type*) extends NonUnitalNormedRing A, StarRing A, CompleteSpace A, + CStarRing A, NormedSpace ℂ A, IsScalarTower ℂ A A, SMulCommClass ℂ A A, StarModule ℂ A where + +/-- The class of non-unital commutative (complex) C⋆-algebras. -/ +class NonUnitalCommCStarAlgebra (A : Type*) extends + NonUnitalNormedCommRing A, NonUnitalCStarAlgebra A + +/-- The class of unital (complex) C⋆-algebras. -/ +class CStarAlgebra (A : Type*) extends NormedRing A, StarRing A, CompleteSpace A, CStarRing A, + NormedAlgebra ℂ A, StarModule ℂ A where + +/-- The class of unital commutative (complex) C⋆-algebras. -/ +class CommCStarAlgebra (A : Type*) extends NormedCommRing A, CStarAlgebra A + +instance (priority := 100) CStarAlgebra.toNonUnitalCStarAlgebra (A : Type*) [CStarAlgebra A] : + NonUnitalCStarAlgebra A where + +instance (priority := 100) CommCStarAlgebra.toNonUnitalCommCStarAlgebra (A : Type*) + [CommCStarAlgebra A] : NonUnitalCommCStarAlgebra A where + +noncomputable instance StarSubalgebra.cstarAlgebra {S A : Type*} [CStarAlgebra A] + [SetLike S A] [SubringClass S A] [SMulMemClass S ℂ A] [StarMemClass S A] + (s : S) [h_closed : IsClosed (s : Set A)] : CStarAlgebra s where + toCompleteSpace := h_closed.completeSpace_coe + norm_mul_self_le x := CStarRing.norm_star_mul_self (x := (x : A)) |>.symm.le + +noncomputable instance StarSubalgebra.commCStarAlgebra {S A : Type*} [CommCStarAlgebra A] + [SetLike S A] [SubringClass S A] [SMulMemClass S ℂ A] [StarMemClass S A] + (s : S) [h_closed : IsClosed (s : Set A)] : CommCStarAlgebra s where + toCompleteSpace := h_closed.completeSpace_coe + norm_mul_self_le x := CStarRing.norm_star_mul_self (x := (x : A)) |>.symm.le + mul_comm _ _ := Subtype.ext <| mul_comm _ _ + +noncomputable instance NonUnitalStarSubalgebra.nonUnitalCStarAlgebra {S A : Type*} + [NonUnitalCStarAlgebra A] [SetLike S A] [NonUnitalSubringClass S A] [SMulMemClass S ℂ A] + [StarMemClass S A] (s : S) [h_closed : IsClosed (s : Set A)] : NonUnitalCStarAlgebra s where + toCompleteSpace := h_closed.completeSpace_coe + norm_mul_self_le x := CStarRing.norm_star_mul_self (x := (x : A)) |>.symm.le + +noncomputable instance NonUnitalStarSubalgebra.nonUnitalCommCStarAlgebra {S A : Type*} + [NonUnitalCommCStarAlgebra A] [SetLike S A] [NonUnitalSubringClass S A] [SMulMemClass S ℂ A] + [StarMemClass S A] (s : S) [h_closed : IsClosed (s : Set A)] : NonUnitalCommCStarAlgebra s where + toCompleteSpace := h_closed.completeSpace_coe + norm_mul_self_le x := CStarRing.norm_star_mul_self (x := (x : A)) |>.symm.le + mul_comm _ _ := Subtype.ext <| mul_comm _ _ + +noncomputable instance : CommCStarAlgebra ℂ where + mul_comm := mul_comm + +section Pi + +variable {ι : Type*} {A : ι → Type*} [Fintype ι] + +instance [(i : ι) → NonUnitalCStarAlgebra (A i)] : NonUnitalCStarAlgebra (Π i, A i) where + +instance [(i : ι) → NonUnitalCommCStarAlgebra (A i)] : NonUnitalCommCStarAlgebra (Π i, A i) where + mul_comm := mul_comm + +noncomputable instance [(i : ι) → CStarAlgebra (A i)] : CStarAlgebra (Π i, A i) where + +noncomputable instance [(i : ι) → CommCStarAlgebra (A i)] : CommCStarAlgebra (Π i, A i) where + mul_comm := mul_comm + +end Pi + +section Prod + +variable {A B : Type*} + +instance [NonUnitalCStarAlgebra A] [NonUnitalCStarAlgebra B] : NonUnitalCStarAlgebra (A × B) where + +instance [NonUnitalCommCStarAlgebra A] [NonUnitalCommCStarAlgebra B] : + NonUnitalCommCStarAlgebra (A × B) where + mul_comm := mul_comm + +noncomputable instance [CStarAlgebra A] [CStarAlgebra B] : CStarAlgebra (A × B) where + +noncomputable instance [CommCStarAlgebra A] [CommCStarAlgebra B] : CommCStarAlgebra (A × B) where + mul_comm := mul_comm + +end Prod diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Basic.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Basic.lean index 073d532f7a479..ddf5ca62ce95f 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Basic.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Basic.lean @@ -48,8 +48,7 @@ open scoped Pointwise ENNReal NNReal ComplexOrder open WeakDual WeakDual.CharacterSpace elementalStarAlgebra -variable {A : Type*} [NormedRing A] [NormedAlgebra ℂ A] -variable [StarRing A] [CStarRing A] [StarModule ℂ A] +variable {A : Type*} [CStarAlgebra A] instance {R A : Type*} [CommRing R] [StarRing R] [NormedRing A] [Algebra R A] [StarRing A] [ContinuousStar A] [StarModule R A] (a : A) [IsStarNormal a] : @@ -57,7 +56,10 @@ instance {R A : Type*} [CommRing R] [StarRing R] [NormedRing A] [Algebra R A] [S { SubringClass.toNormedRing (elementalStarAlgebra R a) with mul_comm := mul_comm } -variable [CompleteSpace A] (a : A) [IsStarNormal a] (S : StarSubalgebra ℂ A) +noncomputable instance (a : A) [IsStarNormal a] : CommCStarAlgebra (elementalStarAlgebra ℂ a) where + mul_comm := mul_comm + +variable (a : A) [IsStarNormal a] /-- The natural map from `characterSpace ℂ (elementalStarAlgebra ℂ x)` to `spectrum ℂ x` given by evaluating `φ` at `x`. This is essentially just evaluation of the `gelfandTransform` of `x`, diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean index 5d998afdff365..1ed5d8debe5ce 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean @@ -154,19 +154,14 @@ end RCLike section Normal -instance IsStarNormal.instContinuousFunctionalCalculus {A : Type*} [NormedRing A] [StarRing A] - [CStarRing A] [CompleteSpace A] [NormedAlgebra ℂ A] [StarModule ℂ A] : +instance IsStarNormal.instContinuousFunctionalCalculus {A : Type*} [CStarAlgebra A] : ContinuousFunctionalCalculus ℂ (IsStarNormal : A → Prop) where predicate_zero := isStarNormal_zero spectrum_nonempty a _ := spectrum.nonempty a exists_cfc_of_predicate a ha := by refine ⟨(elementalStarAlgebra ℂ a).subtype.comp <| continuousFunctionalCalculus a, - ?hom_isClosedEmbedding, ?hom_id, ?hom_map_spectrum, ?predicate_hom⟩ - case hom_isClosedEmbedding => - -- note: Lean should find these for `StarAlgEquiv.isometry`, but it doesn't and so we - -- provide them manually. - have : SMulCommClass ℂ C(σ ℂ a, ℂ) C(σ ℂ a, ℂ) := Algebra.to_smulCommClass (A := C(σ ℂ a, ℂ)) - have : IsScalarTower ℂ C(σ ℂ a, ℂ) C(σ ℂ a, ℂ) := IsScalarTower.right (A := C(σ ℂ a, ℂ)) + ?hom_closedEmbedding, ?hom_id, ?hom_map_spectrum, ?predicate_hom⟩ + case hom_closedEmbedding => exact Isometry.isClosedEmbedding <| isometry_subtype_coe.comp <| StarAlgEquiv.isometry (continuousFunctionalCalculus a) case hom_id => exact congr_arg Subtype.val <| continuousFunctionalCalculus_map_id a @@ -177,10 +172,8 @@ instance IsStarNormal.instContinuousFunctionalCalculus {A : Type*} [NormedRing A AlgEquiv.spectrum_eq (continuousFunctionalCalculus a), ContinuousMap.spectrum_eq_range] case predicate_hom => exact fun f ↦ ⟨by rw [← map_star]; exact Commute.all (star f) f |>.map _⟩ -instance IsStarNormal.instNonUnitalContinuousFunctionalCalculus {A : Type*} [NonUnitalNormedRing A] - [StarRing A] [CStarRing A] [CompleteSpace A] [NormedSpace ℂ A] [IsScalarTower ℂ A A] - [SMulCommClass ℂ A A] [StarModule ℂ A] : - NonUnitalContinuousFunctionalCalculus ℂ (IsStarNormal : A → Prop) := +instance IsStarNormal.instNonUnitalContinuousFunctionalCalculus {A : Type*} + [NonUnitalCStarAlgebra A] : NonUnitalContinuousFunctionalCalculus ℂ (IsStarNormal : A → Prop) := RCLike.nonUnitalContinuousFunctionalCalculus Unitization.isStarNormal_inr end Normal @@ -375,8 +368,7 @@ section SpectrumRestricts open NNReal ENNReal -variable {A : Type*} [NormedRing A] [StarRing A] [CStarRing A] [CompleteSpace A] -variable [NormedAlgebra ℂ A] [StarModule ℂ A] +variable {A : Type*} [CStarAlgebra A] lemma SpectrumRestricts.nnreal_iff_nnnorm {a : A} {t : ℝ≥0} (ha : IsSelfAdjoint a) (ht : ‖a‖₊ ≤ t) : SpectrumRestricts a ContinuousMap.realToNNReal ↔ ‖algebraMap ℝ A t - a‖₊ ≤ t := by @@ -495,9 +487,7 @@ end SpectrumRestricts section NonnegSpectrumClass -variable {A : Type*} [NormedRing A] [CompleteSpace A] -variable [PartialOrder A] [StarRing A] [StarOrderedRing A] [CStarRing A] -variable [NormedAlgebra ℂ A] [StarModule ℂ A] +variable {A : Type*} [CStarAlgebra A] [PartialOrder A] [StarOrderedRing A] instance CStarAlgebra.instNonnegSpectrumClass : NonnegSpectrumClass ℝ A := .of_spectrum_nonneg fun a ha ↦ by @@ -526,8 +516,7 @@ end NonnegSpectrumClass section SpectralOrder -variable (A : Type*) [NormedRing A] [CompleteSpace A] [StarRing A] [CStarRing A] -variable [NormedAlgebra ℂ A] [StarModule ℂ A] +variable (A : Type*) [CStarAlgebra A] /-- The partial order on a unital C⋆-algebra defined by `x ≤ y` if and only if `y - x` is selfadjoint and has nonnegative spectrum. @@ -578,9 +567,7 @@ end SpectralOrder section NonnegSpectrumClass -variable {A : Type*} [NonUnitalNormedRing A] [CompleteSpace A] -variable [PartialOrder A] [StarRing A] [StarOrderedRing A] [CStarRing A] -variable [NormedSpace ℂ A] [IsScalarTower ℂ A A] [SMulCommClass ℂ A A] [StarModule ℂ A] +variable {A : Type*} [NonUnitalCStarAlgebra A] [PartialOrder A] [StarOrderedRing A] instance CStarAlgebra.instNonnegSpectrumClass' : NonnegSpectrumClass ℝ A where quasispectrum_nonneg_of_nonneg a ha := by diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Order.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Order.lean index 08951f013e6d4..75e4270e64c96 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Order.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Order.lean @@ -39,9 +39,7 @@ open scoped NNReal namespace Unitization -variable {A : Type*} [NonUnitalNormedRing A] [CompleteSpace A] - [PartialOrder A] [StarRing A] [StarOrderedRing A] [CStarRing A] [NormedSpace ℂ A] [StarModule ℂ A] - [SMulCommClass ℂ A A] [IsScalarTower ℂ A A] +variable {A : Type*} [NonUnitalCStarAlgebra A] [PartialOrder A] [StarOrderedRing A] instance instPartialOrder : PartialOrder (Unitization ℂ A) := CStarAlgebra.spectralOrder _ @@ -103,8 +101,7 @@ lemma CFC.exists_pos_algebraMap_le_iff {A : Type*} [TopologicalSpace A] [Ring A] section CStar_unital -variable {A : Type*} [NormedRing A] [StarRing A] [CStarRing A] [CompleteSpace A] -variable [NormedAlgebra ℂ A] [StarModule ℂ A] +variable {A : Type*} [CStarAlgebra A] section StarOrderedRing @@ -317,9 +314,7 @@ end CStar_unital section CStar_nonunital -variable {A : Type*} [NonUnitalNormedRing A] [CompleteSpace A] [PartialOrder A] [StarRing A] - [StarOrderedRing A] [CStarRing A] [NormedSpace ℂ A] [StarModule ℂ A] - [SMulCommClass ℂ A A] [IsScalarTower ℂ A A] +variable {A : Type*} [NonUnitalCStarAlgebra A] [PartialOrder A] [StarOrderedRing A] namespace CStarAlgebra diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousLinearMap.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousLinearMap.lean new file mode 100644 index 0000000000000..e1a7aacd09ffa --- /dev/null +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousLinearMap.lean @@ -0,0 +1,17 @@ +/- +Copyright (c) 2024 Jireh Loreaux. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jireh Loreaux +-/ +import Mathlib.Analysis.CStarAlgebra.Classes +import Mathlib.Analysis.InnerProductSpace.Adjoint + +/-! # `E →L[ℂ] E` as a C⋆-algebra + +We place this here because, for reasons related to the import hierarchy, it should not be placed +in earlier files. +-/ + +noncomputable +instance {E : Type*} [NormedAddCommGroup E] [InnerProductSpace ℂ E] [CompleteSpace E] : + CStarAlgebra (E →L[ℂ] E) where diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousMap.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousMap.lean new file mode 100644 index 0000000000000..a0c67afc5074e --- /dev/null +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousMap.lean @@ -0,0 +1,61 @@ +/- +Copyright (c) 2024 Jireh Loreaux. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jireh Loreaux +-/ +import Mathlib.Analysis.CStarAlgebra.Classes +import Mathlib.Topology.ContinuousMap.Compact +import Mathlib.Topology.ContinuousMap.ZeroAtInfty + +/-! # C⋆-algebras of continuous functions + +We place these here because, for reasons related to the import hierarchy, they cannot be placed in +earlier files. +-/ + +variable {α A : Type*} +noncomputable section + +namespace BoundedContinuousFunction + +variable [TopologicalSpace α] + +instance [NonUnitalCStarAlgebra A] : NonUnitalCStarAlgebra (α →ᵇ A) where + +instance [NonUnitalCommCStarAlgebra A] : NonUnitalCommCStarAlgebra (α →ᵇ A) where + mul_comm := mul_comm + +instance [CStarAlgebra A] : CStarAlgebra (α →ᵇ A) where + +instance [CommCStarAlgebra A] : CommCStarAlgebra (α →ᵇ A) where + mul_comm := mul_comm + +end BoundedContinuousFunction + +namespace ContinuousMap + +variable [TopologicalSpace α] [CompactSpace α] + +instance [NonUnitalCStarAlgebra A] : NonUnitalCStarAlgebra C(α, A) where + +instance [NonUnitalCommCStarAlgebra A] : NonUnitalCommCStarAlgebra C(α, A) where + mul_comm := mul_comm + +instance [CStarAlgebra A] : CStarAlgebra C(α, A) where + +instance [CommCStarAlgebra A] : CommCStarAlgebra C(α, A) where + mul_comm := mul_comm + +end ContinuousMap + +namespace ZeroAtInftyContinuousMap + +open ZeroAtInfty + +instance [TopologicalSpace α] [NonUnitalCStarAlgebra A] : NonUnitalCStarAlgebra C₀(α, A) where + +instance [TopologicalSpace α] [NonUnitalCommCStarAlgebra A] : + NonUnitalCommCStarAlgebra C₀(α, A) where + mul_comm := mul_comm + +end ZeroAtInftyContinuousMap diff --git a/Mathlib/Analysis/CStarAlgebra/GelfandDuality.lean b/Mathlib/Analysis/CStarAlgebra/GelfandDuality.lean index a701ede3a9749..88aeab7e08956 100644 --- a/Mathlib/Analysis/CStarAlgebra/GelfandDuality.lean +++ b/Mathlib/Analysis/CStarAlgebra/GelfandDuality.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jireh Loreaux -/ import Mathlib.Analysis.CStarAlgebra.Spectrum +import Mathlib.Analysis.CStarAlgebra.ContinuousMap import Mathlib.Analysis.Normed.Group.Quotient import Mathlib.Analysis.Normed.Algebra.Basic import Mathlib.Topology.ContinuousMap.Units @@ -124,8 +125,7 @@ end ComplexBanachAlgebra section ComplexCStarAlgebra -variable {A : Type*} [NormedCommRing A] [NormedAlgebra ℂ A] [CompleteSpace A] -variable [StarRing A] [CStarRing A] [StarModule ℂ A] +variable {A : Type*} [CommCStarAlgebra A] theorem gelfandTransform_map_star (a : A) : gelfandTransform ℂ A (star a) = star (gelfandTransform ℂ A a) := @@ -258,9 +258,7 @@ V V B --- η B ---> C(characterSpace ℂ B, ℂ) ``` -/ -theorem gelfandStarTransform_naturality {A B : Type*} [NormedCommRing A] [NormedAlgebra ℂ A] - [CompleteSpace A] [StarRing A] [CStarRing A] [StarModule ℂ A] [NormedCommRing B] - [NormedAlgebra ℂ B] [CompleteSpace B] [StarRing B] [CStarRing B] [StarModule ℂ B] +theorem gelfandStarTransform_naturality {A B : Type*} [CommCStarAlgebra A] [CommCStarAlgebra B] (φ : A →⋆ₐ[ℂ] B) : (gelfandStarTransform B : _ →⋆ₐ[ℂ] _).comp φ = (compContinuousMap φ |>.compStarAlgHom' ℂ ℂ).comp (gelfandStarTransform A : _ →⋆ₐ[ℂ] _) := by diff --git a/Mathlib/Analysis/CStarAlgebra/Hom.lean b/Mathlib/Analysis/CStarAlgebra/Hom.lean index 1f062615cf409..86e747a27affb 100644 --- a/Mathlib/Analysis/CStarAlgebra/Hom.lean +++ b/Mathlib/Analysis/CStarAlgebra/Hom.lean @@ -17,11 +17,7 @@ Here we collect properties of C⋆-algebra homomorphisms. -/ open CStarAlgebra in -lemma IsSelfAdjoint.map_spectrum_real {F A B : Type*} - [NormedRing A] [CompleteSpace A] [StarRing A] [CStarRing A] - [NormedAlgebra ℂ A] [StarModule ℂ A] - [NormedRing B] [CompleteSpace B] [StarRing B] [CStarRing B] - [NormedAlgebra ℂ B] [StarModule ℂ B] +lemma IsSelfAdjoint.map_spectrum_real {F A B : Type*} [CStarAlgebra A] [CStarAlgebra B] [FunLike F A B] [AlgHomClass F ℂ A B] [StarHomClass F A B] {a : A} (ha : IsSelfAdjoint a) (φ : F) (hφ : Function.Injective φ) : spectrum ℝ (φ a) = spectrum ℝ a := by @@ -48,11 +44,7 @@ lemma IsSelfAdjoint.map_spectrum_real {F A B : Type*} namespace NonUnitalStarAlgHom -variable {F A B : Type*} -variable [NonUnitalNormedRing A] [CompleteSpace A] [StarRing A] [CStarRing A] -variable [NormedSpace ℂ A] [IsScalarTower ℂ A A] [SMulCommClass ℂ A A] [StarModule ℂ A] -variable [NonUnitalNormedRing B] [CompleteSpace B] [StarRing B] [CStarRing B] -variable [NormedSpace ℂ B] [IsScalarTower ℂ B B] [SMulCommClass ℂ B B] [StarModule ℂ B] +variable {F A B : Type*} [NonUnitalCStarAlgebra A] [NonUnitalCStarAlgebra B] variable [FunLike F A B] [NonUnitalAlgHomClass F ℂ A B] [StarHomClass F A B] open CStarAlgebra Unitization in diff --git a/Mathlib/Analysis/CStarAlgebra/Module/Constructions.lean b/Mathlib/Analysis/CStarAlgebra/Module/Constructions.lean index a2bc174a1c07d..37189e81c7d40 100644 --- a/Mathlib/Analysis/CStarAlgebra/Module/Constructions.lean +++ b/Mathlib/Analysis/CStarAlgebra/Module/Constructions.lean @@ -60,13 +60,13 @@ open CStarModule CStarRing namespace WithCStarModule -variable {A : Type*} [NonUnitalNormedRing A] [StarRing A] [NormedSpace ℂ A] [PartialOrder A] +variable {A : Type*} [NonUnitalCStarAlgebra A] [PartialOrder A] /-! ## A C⋆-algebra as a C⋆-module over itself -/ section Self -variable [CStarRing A] [StarOrderedRing A] [SMulCommClass ℂ A A] +variable [StarOrderedRing A] /-- Reinterpret a C⋆-algebra `A` as a `CStarModule` over itself. -/ instance : CStarModule A A where @@ -112,7 +112,7 @@ lemma prod_norm_le_norm_add (x : C⋆ᵐᵒᵈ (E × F)) : ‖x‖ ≤ ‖x.1‖ _ ≤ ‖x.1‖ ^ 2 + 2 * ‖x.1‖ * ‖x.2‖ + ‖x.2‖ ^ 2 := by gcongr; positivity _ = (‖x.1‖ + ‖x.2‖) ^ 2 := by ring -variable [StarModule ℂ A] [StarOrderedRing A] +variable [StarOrderedRing A] noncomputable instance : CStarModule A (C⋆ᵐᵒᵈ (E × F)) where inner x y := inner x.1 y.1 + inner x.2 y.2 @@ -133,8 +133,6 @@ noncomputable instance : CStarModule A (C⋆ᵐᵒᵈ (E × F)) where lemma prod_inner (x y : C⋆ᵐᵒᵈ (E × F)) : ⟪x, y⟫_A = ⟪x.1, y.1⟫_A + ⟪x.2, y.2⟫_A := rfl -variable [CStarRing A] [SMulCommClass ℂ A A] [IsScalarTower ℂ A A] [CompleteSpace A] - lemma max_le_prod_norm (x : C⋆ᵐᵒᵈ (E × F)) : max ‖x.1‖ ‖x.2‖ ≤ ‖x‖ := by rw [prod_norm] simp only [equiv_fst, norm_eq_sqrt_norm_inner_self (E := E), @@ -214,7 +212,7 @@ lemma pi_norm_le_sum_norm (x : C⋆ᵐᵒᵈ (Π i, E i)) : ‖x‖ ≤ ∑ i, _ = ∑ i, ‖x i‖ ^ 2 := by simp only [norm_sq_eq] _ ≤ (∑ i, ‖x i‖) ^ 2 := sum_sq_le_sq_sum_of_nonneg (fun _ _ ↦ norm_nonneg _) -variable [StarModule ℂ A] [StarOrderedRing A] +variable [StarOrderedRing A] open Finset in noncomputable instance : CStarModule A (C⋆ᵐᵒᵈ (Π i, E i)) where @@ -247,8 +245,6 @@ lemma inner_single_right [DecidableEq ι] (x : C⋆ᵐᵒᵈ (Π i, E i)) {i : rw [Finset.sum_eq_single i] all_goals simp_all -variable [CStarRing A] [SMulCommClass ℂ A A] [IsScalarTower ℂ A A] [CompleteSpace A] - @[simp] lemma norm_single [DecidableEq ι] (i : ι) (y : E i) : ‖equiv _ |>.symm <| Pi.single i y‖ = ‖y‖ := by diff --git a/Mathlib/Analysis/CStarAlgebra/Module/Defs.lean b/Mathlib/Analysis/CStarAlgebra/Module/Defs.lean index 2b601c936ad06..4abfbcecbf696 100644 --- a/Mathlib/Analysis/CStarAlgebra/Module/Defs.lean +++ b/Mathlib/Analysis/CStarAlgebra/Module/Defs.lean @@ -153,8 +153,8 @@ end general section norm -variable {A E : Type*} [NonUnitalNormedRing A] [StarRing A] [PartialOrder A] - [AddCommGroup E] [NormedSpace ℂ A] [Module ℂ E] [SMul Aᵐᵒᵖ E] [Norm E] [CStarModule A E] +variable {A E : Type*} [NonUnitalCStarAlgebra A] [PartialOrder A] [AddCommGroup E] + [Module ℂ E] [SMul Aᵐᵒᵖ E] [Norm E] [CStarModule A E] local notation "⟪" x ", " y "⟫" => inner (𝕜 := A) x y @@ -177,8 +177,6 @@ protected lemma norm_pos {x : E} (hx : x ≠ 0) : 0 < ‖x‖ := by rw [inner_self] at H exact hx H -variable [StarModule ℂ A] - protected lemma norm_zero : ‖(0 : E)‖ = 0 := by simp [norm_eq_sqrt_norm_inner_self] lemma norm_zero_iff (x : E) : ‖x‖ = 0 ↔ x = 0 := @@ -187,12 +185,11 @@ lemma norm_zero_iff (x : E) : ‖x‖ = 0 ↔ x = 0 := end -variable [CStarRing A] [StarOrderedRing A] [StarModule ℂ A] - [IsScalarTower ℂ A A] [SMulCommClass ℂ A A] +variable [StarOrderedRing A] open scoped InnerProductSpace in /-- The C⋆-algebra-valued Cauchy-Schwarz inequality for Hilbert C⋆-modules. -/ -lemma inner_mul_inner_swap_le [CompleteSpace A] {x y : E} : ⟪y, x⟫ * ⟪x, y⟫ ≤ ‖x‖ ^ 2 • ⟪y, y⟫ := by +lemma inner_mul_inner_swap_le {x y : E} : ⟪y, x⟫ * ⟪x, y⟫ ≤ ‖x‖ ^ 2 • ⟪y, y⟫ := by rcases eq_or_ne x 0 with h|h · simp [h, CStarModule.norm_zero (E := E)] · have h₁ : ∀ (a : A), @@ -222,7 +219,7 @@ lemma inner_mul_inner_swap_le [CompleteSpace A] {x y : E} : ⟪y, x⟫ * ⟪x, y open scoped InnerProductSpace in variable (E) in /-- The Cauchy-Schwarz inequality for Hilbert C⋆-modules. -/ -lemma norm_inner_le [CompleteSpace A] {x y : E} : ‖⟪x, y⟫‖ ≤ ‖x‖ * ‖y‖ := by +lemma norm_inner_le {x y : E} : ‖⟪x, y⟫‖ ≤ ‖x‖ * ‖y‖ := by have := calc ‖⟪x, y⟫‖ ^ 2 = ‖⟪y, x⟫ * ⟪x, y⟫‖ := by rw [← star_inner x, CStarRing.norm_star_mul_self, pow_two] _ ≤ ‖‖x‖^ 2 • ⟪y, y⟫‖ := by @@ -236,7 +233,8 @@ lemma norm_inner_le [CompleteSpace A] {x y : E} : ‖⟪x, y⟫‖ ≤ ‖x‖ * refine (pow_le_pow_iff_left (R := ℝ) (norm_nonneg ⟪x, y⟫_A) ?_ (by norm_num)).mp this exact mul_nonneg CStarModule.norm_nonneg CStarModule.norm_nonneg -protected lemma norm_triangle [CompleteSpace A] (x y : E) : ‖x + y‖ ≤ ‖x‖ + ‖y‖ := by +include A in +protected lemma norm_triangle (x y : E) : ‖x + y‖ ≤ ‖x‖ + ‖y‖ := by have h : ‖x + y‖ ^ 2 ≤ (‖x‖ + ‖y‖) ^ 2 := by calc _ ≤ ‖⟪x, x⟫ + ⟪y, x⟫‖ + ‖⟪x, y⟫‖ + ‖⟪y, y⟫‖ := by simp only [norm_eq_sqrt_norm_inner_self, inner_add_right, inner_add_left, ← add_assoc, @@ -250,9 +248,10 @@ protected lemma norm_triangle [CompleteSpace A] (x y : E) : ‖x + y‖ ≤ ‖x refine (pow_le_pow_iff_left CStarModule.norm_nonneg ?_ (by norm_num)).mp h exact add_nonneg CStarModule.norm_nonneg CStarModule.norm_nonneg +include A in /-- This allows us to get `NormedAddCommGroup` and `NormedSpace` instances on `E` via `NormedAddCommGroup.ofCore` and `NormedSpace.ofCore`. -/ -lemma normedSpaceCore [CompleteSpace A] : NormedSpace.Core ℂ E where +lemma normedSpaceCore : NormedSpace.Core ℂ E where norm_nonneg _ := CStarModule.norm_nonneg norm_eq_zero_iff x := norm_zero_iff x norm_smul c x := by simp [norm_eq_sqrt_norm_inner_self, norm_smul, ← mul_assoc] @@ -260,11 +259,11 @@ lemma normedSpaceCore [CompleteSpace A] : NormedSpace.Core ℂ E where /-- This is not listed as an instance because we often want to replace the topology, uniformity and bornology instead of inheriting them from the norm. -/ -abbrev normedAddCommGroup [CompleteSpace A] : NormedAddCommGroup E := +abbrev normedAddCommGroup : NormedAddCommGroup E := NormedAddCommGroup.ofCore CStarModule.normedSpaceCore open scoped InnerProductSpace in -lemma norm_eq_csSup [CompleteSpace A] (v : E) : +lemma norm_eq_csSup (v : E) : ‖v‖ = sSup { ‖⟪w, v⟫_A‖ | (w : E) (_ : ‖w‖ ≤ 1) } := by let instNACG : NormedAddCommGroup E := NormedAddCommGroup.ofCore normedSpaceCore let instNS : NormedSpace ℂ E := .ofCore normedSpaceCore @@ -286,10 +285,8 @@ open scoped InnerProductSpace `NormedAddCommGroup` and `NormedSpace` instances via `CStarModule.normedSpaceCore`, especially by using `NormedAddCommGroup.ofCoreReplaceAll` and `NormedSpace.ofCore`. See `Analysis.CStarAlgebra.Module.Constructions` for examples. -/ -variable {A E : Type*} [NonUnitalNormedRing A] [StarRing A] [CStarRing A] [PartialOrder A] - [StarOrderedRing A] [NormedSpace ℂ A] [SMul Aᵐᵒᵖ E] [CompleteSpace A] - [NormedAddCommGroup E] [NormedSpace ℂ E] [StarModule ℂ A] [CStarModule A E] [IsScalarTower ℂ A A] - [SMulCommClass ℂ A A] +variable {A E : Type*} [NonUnitalCStarAlgebra A] [PartialOrder A] [StarOrderedRing A] [SMul Aᵐᵒᵖ E] + [NormedAddCommGroup E] [NormedSpace ℂ E] [CStarModule A E] /-- The function `⟨x, y⟩ ↦ ⟪x, y⟫` bundled as a continuous sesquilinear map. -/ noncomputable def innerSL : E →L⋆[ℂ] E →L[ℂ] A := diff --git a/Mathlib/Analysis/CStarAlgebra/Multiplier.lean b/Mathlib/Analysis/CStarAlgebra/Multiplier.lean index d5f3b758aada7..2bba23247166b 100644 --- a/Mathlib/Analysis/CStarAlgebra/Multiplier.lean +++ b/Mathlib/Analysis/CStarAlgebra/Multiplier.lean @@ -5,6 +5,7 @@ Authors: Jireh Loreaux, Jon Bannon -/ import Mathlib.Analysis.NormedSpace.OperatorNorm.Completeness import Mathlib.Analysis.CStarAlgebra.Unitization +import Mathlib.Analysis.CStarAlgebra.Classes import Mathlib.Analysis.SpecialFunctions.Pow.NNReal /-! @@ -667,4 +668,6 @@ instance instCStarRing : CStarRing 𝓜(𝕜, A) where end DenselyNormed +noncomputable instance {A : Type*} [NonUnitalCStarAlgebra A] : CStarAlgebra 𝓜(ℂ, A) where + end DoubleCentralizer diff --git a/Mathlib/Analysis/CStarAlgebra/Spectrum.lean b/Mathlib/Analysis/CStarAlgebra/Spectrum.lean index 65c73492f7890..ad5459084d830 100644 --- a/Mathlib/Analysis/CStarAlgebra/Spectrum.lean +++ b/Mathlib/Analysis/CStarAlgebra/Spectrum.lean @@ -90,8 +90,7 @@ section ComplexScalars open Complex -variable {A : Type*} [NormedRing A] [NormedAlgebra ℂ A] [CompleteSpace A] [StarRing A] - [CStarRing A] +variable {A : Type*} [CStarAlgebra A] local notation "↑ₐ" => algebraMap ℂ A @@ -221,11 +220,7 @@ end ComplexScalars namespace NonUnitalStarAlgHom -variable {F A B : Type*} -variable [NonUnitalNormedRing A] [CompleteSpace A] [StarRing A] [CStarRing A] -variable [NormedSpace ℂ A] [IsScalarTower ℂ A A] [SMulCommClass ℂ A A] [StarModule ℂ A] -variable [NonUnitalNormedRing B] [CompleteSpace B] [StarRing B] [CStarRing B] -variable [NormedSpace ℂ B] [IsScalarTower ℂ B B] [SMulCommClass ℂ B B] [StarModule ℂ B] +variable {F A B : Type*} [NonUnitalCStarAlgebra A] [NonUnitalCStarAlgebra B] variable [FunLike F A B] [NonUnitalAlgHomClass F ℂ A B] [StarHomClass F A B] open Unitization @@ -240,7 +235,6 @@ lemma nnnorm_apply_le (φ : F) (a : A) : ‖φ a‖₊ ≤ ‖a‖₊ := by exact this <| .star_mul_self x intro s hs suffices this : spectralRadius ℂ (ψ s) ≤ spectralRadius ℂ s by - -- changing the order of `rw`s below runs into https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/weird.20type.20class.20synthesis.20error/near/421224482 rwa [(hs.map ψ).spectralRadius_eq_nnnorm, hs.spectralRadius_eq_nnnorm, coe_le_coe] at this exact iSup_le_iSup_of_subset (AlgHom.spectrum_apply_subset ψ s) @@ -263,10 +257,7 @@ end NonUnitalStarAlgHom namespace StarAlgEquiv -variable {F A B : Type*} [NonUnitalNormedRing A] [NormedSpace ℂ A] [SMulCommClass ℂ A A] -variable [IsScalarTower ℂ A A] [CompleteSpace A] [StarRing A] [CStarRing A] [StarModule ℂ A] -variable [NonUnitalNormedRing B] [NormedSpace ℂ B] [SMulCommClass ℂ B B] [IsScalarTower ℂ B B] -variable [CompleteSpace B] [StarRing B] [CStarRing B] [StarModule ℂ B] [EquivLike F A B] +variable {F A B : Type*} [NonUnitalCStarAlgebra A] [NonUnitalCStarAlgebra B] [EquivLike F A B] variable [NonUnitalAlgEquivClass F ℂ A B] [StarHomClass F A B] lemma nnnorm_map (φ : F) (a : A) : ‖φ a‖₊ = ‖a‖₊ := @@ -289,8 +280,7 @@ open ContinuousMap Complex open scoped ComplexStarModule -variable {F A : Type*} [NormedRing A] [NormedAlgebra ℂ A] [CompleteSpace A] [StarRing A] - [CStarRing A] [StarModule ℂ A] [FunLike F A ℂ] [hF : AlgHomClass F ℂ A ℂ] +variable {F A : Type*} [CStarAlgebra A] [FunLike F A ℂ] [hF : AlgHomClass F ℂ A ℂ] /-- This instance is provided instead of `StarHomClass` to avoid type class inference loops. See note [lower instance priority] -/ diff --git a/Mathlib/Analysis/CStarAlgebra/Unitization.lean b/Mathlib/Analysis/CStarAlgebra/Unitization.lean index 7a5bdee2eef6e..0c5534a194ce1 100644 --- a/Mathlib/Analysis/CStarAlgebra/Unitization.lean +++ b/Mathlib/Analysis/CStarAlgebra/Unitization.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Jireh Loreaux. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jireh Loreaux -/ -import Mathlib.Analysis.CStarAlgebra.Basic +import Mathlib.Analysis.CStarAlgebra.Classes import Mathlib.Analysis.Normed.Algebra.Unitization /-! # The minimal unitization of a C⋆-algebra @@ -30,10 +30,13 @@ variable [NormedSpace 𝕜 E] [IsScalarTower 𝕜 E E] [SMulCommClass 𝕜 E E] lemma opNorm_mul_flip_apply (a : E) : ‖(mul 𝕜 E).flip a‖ = ‖a‖ := by refine le_antisymm (opNorm_le_bound _ (norm_nonneg _) fun b => by simpa only [mul_comm] using norm_mul_le b a) ?_ - suffices ‖mul 𝕜 E (star a)‖ ≤ ‖(mul 𝕜 E).flip a‖ by simpa using this + suffices ‖mul 𝕜 E (star a)‖ ≤ ‖(mul 𝕜 E).flip a‖ by + simpa only [ge_iff_le, opNorm_mul_apply, norm_star] using this refine opNorm_le_bound _ (norm_nonneg _) fun b => ?_ - calc ‖mul 𝕜 E (star a) b‖ = ‖(mul 𝕜 E).flip a (star b)‖ := by simpa using norm_star (star b * a) - _ ≤ ‖(mul 𝕜 E).flip a‖ * ‖b‖ := by simpa using le_opNorm ((mul 𝕜 E).flip a) (star b) + calc ‖mul 𝕜 E (star a) b‖ = ‖(mul 𝕜 E).flip a (star b)‖ := by + simpa only [mul_apply', flip_apply, star_mul, star_star] using norm_star (star b * a) + _ ≤ ‖(mul 𝕜 E).flip a‖ * ‖b‖ := by + simpa only [flip_apply, mul_apply', norm_star] using le_opNorm ((mul 𝕜 E).flip a) (star b) @[deprecated (since := "2024-02-02")] alias op_norm_mul_flip_apply := opNorm_mul_flip_apply @@ -171,4 +174,15 @@ instance Unitization.instCStarRing : CStarRing (Unitization 𝕜 E) where · replace h := (not_le.mp h).le rw [sq, sq, sup_eq_left.mpr h, sup_eq_left.mpr (mul_self_le_mul_self (norm_nonneg _) h)] +/-- The minimal unitization (over `ℂ`) of a C⋆-algebra, equipped with the C⋆-norm. When `A` is +unital, `A⁺¹ ≃⋆ₐ[ℂ] (ℂ × A)`. -/ +scoped[CStarAlgebra] postfix:max "⁺¹" => Unitization ℂ + +noncomputable instance Unitization.instCStarAlgebra {A : Type*} [NonUnitalCStarAlgebra A] : + CStarAlgebra (Unitization ℂ A) where + +noncomputable instance Unitization.instCommCStarAlgebra {A : Type*} [NonUnitalCommCStarAlgebra A] : + CommCStarAlgebra (Unitization ℂ A) where + mul_comm := mul_comm + end CStarProperty diff --git a/Mathlib/Analysis/CStarAlgebra/lpSpace.lean b/Mathlib/Analysis/CStarAlgebra/lpSpace.lean new file mode 100644 index 0000000000000..26dd3dc11170c --- /dev/null +++ b/Mathlib/Analysis/CStarAlgebra/lpSpace.lean @@ -0,0 +1,35 @@ +/- +Copyright (c) 2024 Jireh Loreaux. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jireh Loreaux +-/ +import Mathlib.Analysis.CStarAlgebra.Classes +import Mathlib.Analysis.Normed.Lp.lpSpace + +/-! # `lp ∞ A` as a C⋆-algebra + +We place these here because, for reasons related to the import hierarchy, they should not be placed +in earlier files. +-/ +open scoped ENNReal + +noncomputable section + +variable {I : Type*} {A : I → Type*} + +instance [∀ i, NonUnitalCStarAlgebra (A i)] : NonUnitalCStarAlgebra (lp A ∞) where + +instance [∀ i, NonUnitalCommCStarAlgebra (A i)] : NonUnitalCommCStarAlgebra (lp A ∞) where + mul_comm := mul_comm + +-- it's slightly weird that we need the `Nontrivial` instance here +-- it's because we have no way to say that `‖(1 : A i)‖` is uniformly bounded as a type class +-- aside from `∀ i, NormOneClass (A i)`, this holds automatically for C⋆-algebras though. +instance [∀ i, Nontrivial (A i)] [∀ i, CStarAlgebra (A i)] : NormedRing (lp A ∞) where + dist_eq := dist_eq_norm + norm_mul := norm_mul_le + +instance [∀ i, Nontrivial (A i)] [∀ i, CommCStarAlgebra (A i)] : CommCStarAlgebra (lp A ∞) where + mul_comm := mul_comm + +end diff --git a/Mathlib/Analysis/InnerProductSpace/Adjoint.lean b/Mathlib/Analysis/InnerProductSpace/Adjoint.lean index b6803862c58c6..0537bba885e12 100644 --- a/Mathlib/Analysis/InnerProductSpace/Adjoint.lean +++ b/Mathlib/Analysis/InnerProductSpace/Adjoint.lean @@ -215,6 +215,8 @@ theorem norm_adjoint_comp_self (A : E →L[𝕜] F) : simp_rw [mul_assoc, Real.sqrt_mul (norm_nonneg _) (‖x‖ * ‖x‖), Real.sqrt_mul_self (norm_nonneg x)] +/-- The C⋆-algebra instance when `𝕜 := ℂ` can be found in +`Analysis.CStarAlgebra.ContinuousLinearMap`. -/ instance : CStarRing (E →L[𝕜] E) where norm_mul_self_le x := le_of_eq <| Eq.symm <| norm_adjoint_comp_self x diff --git a/Mathlib/Analysis/InnerProductSpace/StarOrder.lean b/Mathlib/Analysis/InnerProductSpace/StarOrder.lean index cbf503358c95e..371413581895d 100644 --- a/Mathlib/Analysis/InnerProductSpace/StarOrder.lean +++ b/Mathlib/Analysis/InnerProductSpace/StarOrder.lean @@ -5,6 +5,7 @@ Authors: Jireh Loreaux -/ import Mathlib.Analysis.InnerProductSpace.Positive import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Instances +import Mathlib.Analysis.CStarAlgebra.ContinuousLinearMap /-! # Continuous linear maps on a Hilbert space are a `StarOrderedRing` diff --git a/Mathlib/Analysis/Normed/Lp/lpSpace.lean b/Mathlib/Analysis/Normed/Lp/lpSpace.lean index a3ebaed3c2b34..e9bce7f9e90db 100644 --- a/Mathlib/Analysis/Normed/Lp/lpSpace.lean +++ b/Mathlib/Analysis/Normed/Lp/lpSpace.lean @@ -736,6 +736,10 @@ instance nonUnitalNormedRing : NonUnitalNormedRing (lp B ∞) := mul_le_mul (lp.norm_apply_le_norm ENNReal.top_ne_zero f i) (lp.norm_apply_le_norm ENNReal.top_ne_zero g i) (norm_nonneg _) (norm_nonneg _) } +instance nonUnitalNormedCommRing {B : I → Type*} [∀ i, NonUnitalNormedCommRing (B i)] : + NonUnitalNormedCommRing (lp B ∞) where + mul_comm _ _ := ext <| mul_comm .. + -- we also want a `NonUnitalNormedCommRing` instance, but this has to wait for mathlib3 #13719 instance infty_isScalarTower {𝕜} [NormedRing 𝕜] [∀ i, Module 𝕜 (B i)] [∀ i, BoundedSMul 𝕜 (B i)] [∀ i, IsScalarTower 𝕜 (B i) (B i)] : IsScalarTower 𝕜 (lp B ∞) (lp B ∞) := @@ -840,12 +844,8 @@ section NormedCommRing variable {I : Type*} {B : I → Type*} [∀ i, NormedCommRing (B i)] [∀ i, NormOneClass (B i)] -instance inftyCommRing : CommRing (lp B ∞) := - { lp.inftyRing with - mul_comm := fun f g => by ext; simp only [lp.infty_coeFn_mul, Pi.mul_apply, mul_comm] } - -instance inftyNormedCommRing : NormedCommRing (lp B ∞) := - { lp.inftyCommRing, lp.inftyNormedRing with } +instance inftyNormedCommRing : NormedCommRing (lp B ∞) where + mul_comm := mul_comm end NormedCommRing diff --git a/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/ExpLog.lean b/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/ExpLog.lean index 65c353f847076..27a9142d793a9 100644 --- a/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/ExpLog.lean +++ b/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/ExpLog.lean @@ -95,8 +95,7 @@ end RealNormed section ComplexNormed variable {A : Type*} {p : A → Prop} [NormedRing A] [StarRing A] - [TopologicalRing A] [NormedAlgebra ℂ A] [CompleteSpace A] - [ContinuousFunctionalCalculus ℂ p] + [NormedAlgebra ℂ A] [CompleteSpace A] [ContinuousFunctionalCalculus ℂ p] lemma complex_exp_eq_normedSpace_exp {a : A} (ha : p a := by cfc_tac) : cfc Complex.exp a = exp ℂ a := diff --git a/Mathlib/Analysis/VonNeumannAlgebra/Basic.lean b/Mathlib/Analysis/VonNeumannAlgebra/Basic.lean index 0d05481d7f037..4bf4e8678ab11 100644 --- a/Mathlib/Analysis/VonNeumannAlgebra/Basic.lean +++ b/Mathlib/Analysis/VonNeumannAlgebra/Basic.lean @@ -3,9 +3,8 @@ Copyright (c) 2022 Kim Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison -/ +import Mathlib.Analysis.CStarAlgebra.ContinuousLinearMap import Mathlib.Analysis.Normed.Module.Dual -import Mathlib.Analysis.Complex.Basic -import Mathlib.Analysis.InnerProductSpace.Adjoint /-! # Von Neumann algebras @@ -39,8 +38,7 @@ One the other hand, not picking one means that the weak-* topology (which depends on a choice of predual) must be defined using the choice, and we may be unhappy with the resulting opaqueness of the definition. -/ -class WStarAlgebra (M : Type u) [NormedRing M] [StarRing M] [CStarRing M] [Module ℂ M] - [NormedAlgebra ℂ M] [StarModule ℂ M] : Prop where +class WStarAlgebra (M : Type u) [CStarAlgebra M] : Prop where /-- There is a Banach space `X` whose dual is isometrically (conjugate-linearly) isomorphic to the `WStarAlgebra`. -/ exists_predual : diff --git a/Mathlib/Topology/ContinuousMap/Compact.lean b/Mathlib/Topology/ContinuousMap/Compact.lean index e5b573631c838..24fb53d8c00e8 100644 --- a/Mathlib/Topology/ContinuousMap/Compact.lean +++ b/Mathlib/Topology/ContinuousMap/Compact.lean @@ -527,9 +527,9 @@ end NormedSpace section CStarRing variable {α : Type*} {β : Type*} -variable [TopologicalSpace α] [NonUnitalNormedRing β] [StarRing β] +variable [TopologicalSpace α] [CompactSpace α] -instance [CompactSpace α] [CStarRing β] : CStarRing C(α, β) where +instance [NonUnitalNormedRing β] [StarRing β] [CStarRing β] : CStarRing C(α, β) where norm_mul_self_le f := by rw [← sq, ← Real.le_sqrt (norm_nonneg _) (norm_nonneg _), ContinuousMap.norm_le _ (Real.sqrt_nonneg _)] From ff832ebed6cd8b3c62702c0f4b3b34251a45f792 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Thu, 24 Oct 2024 15:52:26 +0000 Subject: [PATCH 378/505] =?UTF-8?q?chore:=20use=20alias=20for=20duplicate?= =?UTF-8?q?=20lemma=20MvPolynomial.eval=E2=82=82Hom=5FC=5Fleft=20(#17263)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- Mathlib/Algebra/MvPolynomial/Monad.lean | 3 +-- 1 file changed, 1 insertion(+), 2 deletions(-) diff --git a/Mathlib/Algebra/MvPolynomial/Monad.lean b/Mathlib/Algebra/MvPolynomial/Monad.lean index ac7de02a758d8..aa9516b1d70fd 100644 --- a/Mathlib/Algebra/MvPolynomial/Monad.lean +++ b/Mathlib/Algebra/MvPolynomial/Monad.lean @@ -268,8 +268,7 @@ theorem aeval_bind₂ [Algebra S T] (f : σ → T) (g : R →+* MvPolynomial σ aeval f (bind₂ g φ) = eval₂Hom ((↑(aeval f : _ →ₐ[S] _) : _ →+* _).comp g) f φ := eval₂Hom_bind₂ _ _ _ _ -theorem eval₂Hom_C_left (f : σ → MvPolynomial τ R) : eval₂Hom C f = bind₁ f := - rfl +alias eval₂Hom_C_left := eval₂Hom_C_eq_bind₁ theorem bind₁_monomial (f : σ → MvPolynomial τ R) (d : σ →₀ ℕ) (r : R) : bind₁ f (monomial d r) = C r * ∏ i ∈ d.support, f i ^ d i := by From 8d463aee83e8cdc415b744069803d0c6c491c383 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Thu, 24 Oct 2024 16:05:10 +0000 Subject: [PATCH 379/505] feat(RingTheory/Kaehler): kaehler differential module under base change (#18153) Port of https://github.com/leanprover-community/mathlib3/pull/17198 Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- Mathlib.lean | 1 + Mathlib/RingTheory/Kaehler/TensorProduct.lean | 233 ++++++++++++++++++ 2 files changed, 234 insertions(+) create mode 100644 Mathlib/RingTheory/Kaehler/TensorProduct.lean diff --git a/Mathlib.lean b/Mathlib.lean index 34a16720097ff..2ee5eb975f465 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4050,6 +4050,7 @@ import Mathlib.RingTheory.JacobsonIdeal import Mathlib.RingTheory.Kaehler.Basic import Mathlib.RingTheory.Kaehler.CotangentComplex import Mathlib.RingTheory.Kaehler.Polynomial +import Mathlib.RingTheory.Kaehler.TensorProduct import Mathlib.RingTheory.KrullDimension.Basic import Mathlib.RingTheory.KrullDimension.Field import Mathlib.RingTheory.LaurentSeries diff --git a/Mathlib/RingTheory/Kaehler/TensorProduct.lean b/Mathlib/RingTheory/Kaehler/TensorProduct.lean new file mode 100644 index 0000000000000..6999913ad12a4 --- /dev/null +++ b/Mathlib/RingTheory/Kaehler/TensorProduct.lean @@ -0,0 +1,233 @@ +/- +Copyright (c) 2024 Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Andrew Yang +-/ +import Mathlib.RingTheory.Kaehler.Basic +import Mathlib.RingTheory.Localization.BaseChange + +/-! +# Kaehler differential module under base change + +## Main results +- `KaehlerDifferential.tensorKaehlerEquiv`: `(S ⊗[R] Ω[A⁄R]) ≃ₗ[S] Ω[B⁄S]` for `B = S ⊗[R] A`. +- `KaehlerDifferential.isLocalizedModule_of_isLocalizedModule`: + `Ω[Aₚ/Rₚ]` is the localization of `Ω[A/R]` at `p`. + +-/ + +variable (R S A B : Type*) [CommRing R] [CommRing S] [Algebra R S] [CommRing A] [CommRing B] +variable [Algebra R A] [Algebra R B] +variable [Algebra A B] [Algebra S B] [IsScalarTower R A B] [IsScalarTower R S B] + +open TensorProduct + +attribute [local instance] SMulCommClass.of_commMonoid + +namespace KaehlerDifferential + +/-- (Implementation). `A`-action on `S ⊗[R] Ω[A⁄R]`. -/ +noncomputable +abbrev mulActionBaseChange : + MulAction A (S ⊗[R] Ω[A⁄R]) := (TensorProduct.comm R S (Ω[A⁄R])).toEquiv.mulAction A + +attribute [local instance] mulActionBaseChange + +@[simp] +lemma mulActionBaseChange_smul_tmul (a : A) (s : S) (x : Ω[A⁄R]) : + a • (s ⊗ₜ[R] x) = s ⊗ₜ (a • x) := rfl + +@[local simp] +lemma mulActionBaseChange_smul_zero (a : A) : + a • (0 : S ⊗[R] Ω[A⁄R]) = 0 := by + rw [← zero_tmul _ (0 : Ω[A⁄R]), mulActionBaseChange_smul_tmul, smul_zero] + +@[local simp] +lemma mulActionBaseChange_smul_add (a : A) (x y : S ⊗[R] Ω[A⁄R]) : + a • (x + y) = a • x + a • y := by + show (TensorProduct.comm R S (Ω[A⁄R])).symm (a • (TensorProduct.comm R S (Ω[A⁄R])) (x + y)) = _ + rw [map_add, smul_add, map_add] + rfl + +/-- (Implementation). `A`-module structure on `S ⊗[R] Ω[A⁄R]`. -/ +noncomputable +abbrev moduleBaseChange : + Module A (S ⊗[R] Ω[A⁄R]) where + __ := (TensorProduct.comm R S (Ω[A⁄R])).toEquiv.mulAction A + add_smul r s x := by induction x <;> simp [add_smul, tmul_add, *, add_add_add_comm] + zero_smul x := by induction x <;> simp [*] + smul_zero := by simp + smul_add := by simp + +attribute [local instance] moduleBaseChange + +instance : IsScalarTower R A (S ⊗[R] Ω[A⁄R]) := by + apply IsScalarTower.of_algebraMap_smul + intro r x + induction x + · simp only [smul_zero] + · rw [mulActionBaseChange_smul_tmul, algebraMap_smul, tmul_smul] + · simp only [smul_add, *] + +instance : SMulCommClass S A (S ⊗[R] Ω[A⁄R]) where + smul_comm s a x := by + induction x + · simp only [smul_zero] + · rw [mulActionBaseChange_smul_tmul, smul_tmul', smul_tmul', mulActionBaseChange_smul_tmul] + · simp only [smul_add, *] + +instance : SMulCommClass A S (S ⊗[R] Ω[A⁄R]) where + smul_comm s a x := by rw [← smul_comm] + +/-- (Implementation). `B = S ⊗[R] A`-module structure on `S ⊗[R] Ω[A⁄R]`. -/ +@[reducible] noncomputable +def moduleBaseChange' [Algebra.IsPushout R S A B] : + Module B (S ⊗[R] Ω[A⁄R]) := + Module.compHom _ (Algebra.pushoutDesc B (Algebra.lsmul R (A := S) S (S ⊗[R] Ω[A⁄R])) + (Algebra.lsmul R (A := A) _ _) (LinearMap.ext <| smul_comm · ·)).toRingHom + +attribute [local instance] moduleBaseChange' + +instance [Algebra.IsPushout R S A B] : + IsScalarTower A B (S ⊗[R] Ω[A⁄R]) := by + apply IsScalarTower.of_algebraMap_smul + intro r x + show (Algebra.pushoutDesc B (Algebra.lsmul R (A := S) S (S ⊗[R] Ω[A⁄R])) + (Algebra.lsmul R (A := A) _ _) (LinearMap.ext <| smul_comm · ·) + (algebraMap A B r)) • x = r • x + simp only [Algebra.pushoutDesc_right, LinearMap.smul_def, Algebra.lsmul_coe] + +instance [Algebra.IsPushout R S A B] : + IsScalarTower S B (S ⊗[R] Ω[A⁄R]) := by + apply IsScalarTower.of_algebraMap_smul + intro r x + show (Algebra.pushoutDesc B (Algebra.lsmul R (A := S) S (S ⊗[R] Ω[A⁄R])) + (Algebra.lsmul R (A := A) _ _) (LinearMap.ext <| smul_comm · ·) + (algebraMap S B r)) • x = r • x + simp only [Algebra.pushoutDesc_left, LinearMap.smul_def, Algebra.lsmul_coe] + +lemma map_liftBaseChange_smul [h : Algebra.IsPushout R S A B] (b : B) (x) : + ((map R S A B).restrictScalars R).liftBaseChange S (b • x) = + b • ((map R S A B).restrictScalars R).liftBaseChange S x := by + induction b using h.1.inductionOn with + | h₁ => simp only [zero_smul, map_zero] + | h₃ s b e => rw [smul_assoc, map_smul, e, smul_assoc] + | h₄ b₁ b₂ e₁ e₂ => simp only [map_add, e₁, e₂, add_smul] + | h₂ a => + induction x + · simp only [smul_zero, map_zero] + · simp [smul_comm] + · simp only [map_add, smul_add, *] + +/-- (Implementation). +The `S`-derivation `B = S ⊗[R] A` to `S ⊗[R] Ω[A⁄R]` sending `a ⊗ b` to `a ⊗ d b`. -/ +noncomputable +def derivationTensorProduct [h : Algebra.IsPushout R S A B] : + Derivation S B (S ⊗[R] Ω[A⁄R]) where + __ := h.out.lift ((TensorProduct.mk R S (Ω[A⁄R]) 1).comp (D R A).toLinearMap) + map_one_eq_zero' := by + rw [← (algebraMap A B).map_one] + refine (h.out.lift_eq _ _).trans ?_ + dsimp + rw [Derivation.map_one_eq_zero, TensorProduct.tmul_zero] + leibniz' a b := by + dsimp + induction a using h.out.inductionOn with + | h₁ => rw [map_zero, zero_smul, smul_zero, zero_add, zero_mul, map_zero] + | h₃ x y e => + rw [smul_mul_assoc, map_smul, e, map_smul, smul_add, + smul_comm x b, smul_assoc] + | h₄ b₁ b₂ e₁ e₂ => simp only [add_mul, add_smul, map_add, e₁, e₂, smul_add, add_add_add_comm] + | h₂ z => + dsimp + induction b using h.out.inductionOn with + | h₁ => rw [map_zero, zero_smul, smul_zero, zero_add, mul_zero, map_zero] + | h₂ => + simp only [AlgHom.toLinearMap_apply, IsScalarTower.coe_toAlgHom', + algebraMap_smul, ← map_mul] + erw [h.out.lift_eq, h.out.lift_eq, h.out.lift_eq] + simp only [LinearMap.coe_comp, Derivation.coeFn_coe, Function.comp_apply, + Derivation.leibniz, mk_apply, mulActionBaseChange_smul_tmul, TensorProduct.tmul_add] + | h₃ _ _ e => + rw [mul_comm, smul_mul_assoc, map_smul, mul_comm, e, + map_smul, smul_add, smul_comm, smul_assoc] + | h₄ _ _ e₁ e₂ => simp only [mul_add, add_smul, map_add, e₁, e₂, smul_add, add_add_add_comm] + +lemma derivationTensorProduct_algebraMap [Algebra.IsPushout R S A B] (x) : + derivationTensorProduct R S A B (algebraMap A B x) = + 1 ⊗ₜ D _ _ x := +IsBaseChange.lift_eq _ _ _ + +lemma tensorKaehlerEquiv_left_inv [Algebra.IsPushout R S A B] : + ((derivationTensorProduct R S A B).liftKaehlerDifferential.restrictScalars S).comp + (((map R S A B).restrictScalars R).liftBaseChange S) = LinearMap.id := by + refine LinearMap.restrictScalars_injective R ?_ + apply TensorProduct.ext' + intro x y + obtain ⟨y, rfl⟩ := tensorProductTo_surjective _ _ y + induction y + · simp only [map_zero, TensorProduct.tmul_zero] + · simp only [LinearMap.restrictScalars_comp, Derivation.tensorProductTo_tmul, LinearMap.coe_comp, + LinearMap.coe_restrictScalars, Function.comp_apply, LinearMap.liftBaseChange_tmul, map_smul, + map_D, LinearMap.map_smul_of_tower, Derivation.liftKaehlerDifferential_comp_D, + LinearMap.id_coe, id_eq, derivationTensorProduct_algebraMap] + rw [smul_comm, TensorProduct.smul_tmul', smul_eq_mul, mul_one] + rfl + · simp only [map_add, TensorProduct.tmul_add, *] + +/-- The canonical isomorphism `(S ⊗[R] Ω[A⁄R]) ≃ₗ[S] Ω[B⁄S]` for `B = S ⊗[R] A`. -/ +@[simps! symm_apply] noncomputable +def tensorKaehlerEquiv [h : Algebra.IsPushout R S A B] : + (S ⊗[R] Ω[A⁄R]) ≃ₗ[S] Ω[B⁄S] where + __ := ((map R S A B).restrictScalars R).liftBaseChange S + invFun := (derivationTensorProduct R S A B).liftKaehlerDifferential + left_inv := LinearMap.congr_fun (tensorKaehlerEquiv_left_inv R S A B) + right_inv x := by + obtain ⟨x, rfl⟩ := tensorProductTo_surjective _ _ x + dsimp + induction x with + | zero => simp + | add x y e₁ e₂ => simp only [map_add, e₁, e₂] + | tmul x y => + dsimp + simp only [Derivation.tensorProductTo_tmul, LinearMap.map_smul, + Derivation.liftKaehlerDifferential_comp_D, map_liftBaseChange_smul] + induction y using h.1.inductionOn + · simp only [map_zero, smul_zero] + · simp only [AlgHom.toLinearMap_apply, IsScalarTower.coe_toAlgHom', + derivationTensorProduct_algebraMap, LinearMap.liftBaseChange_tmul, + LinearMap.coe_restrictScalars, map_D, one_smul] + · simp only [Derivation.map_smul, LinearMap.map_smul, *, smul_comm x] + · simp only [map_add, smul_add, *] + +@[simp] +lemma tensorKaehlerEquiv_tmul [Algebra.IsPushout R S A B] (a b) : + tensorKaehlerEquiv R S A B (a ⊗ₜ b) = a • map R S A B b := + LinearMap.liftBaseChange_tmul _ _ _ _ + +/-- +If `B` is the tensor product of `S` and `A` over `R`, +then `Ω[B⁄S]` is the base change of `Ω[A⁄R]` along `R → S`. +-/ +lemma isBaseChange [h : Algebra.IsPushout R S A B] : + IsBaseChange S ((map R S A B).restrictScalars R) := by + convert (TensorProduct.isBaseChange R (Ω[A⁄R]) S).comp + (IsBaseChange.ofEquiv (tensorKaehlerEquiv R S A B)) + refine LinearMap.ext fun x ↦ ?_ + simp only [LinearMap.coe_restrictScalars, LinearMap.coe_comp, LinearEquiv.coe_coe, + Function.comp_apply, mk_apply, tensorKaehlerEquiv_tmul, one_smul] + +instance isLocalizedModule (p : Submonoid R) [IsLocalization p S] + [IsLocalization (Algebra.algebraMapSubmonoid A p) B] : + IsLocalizedModule p ((map R S A B).restrictScalars R) := + have := (Algebra.isPushout_of_isLocalization p S A B).symm + (isLocalizedModule_iff_isBaseChange p S _).mpr (isBaseChange R S A B) + +instance isLocalizedModule_of_isLocalizedModule (p : Submonoid R) [IsLocalization p S] + [IsLocalizedModule p (IsScalarTower.toAlgHom R A B).toLinearMap] : + IsLocalizedModule p ((map R S A B).restrictScalars R) := + have : IsLocalization (Algebra.algebraMapSubmonoid A p) B := + isLocalizedModule_iff_isLocalization.mp inferInstance + inferInstance + +end KaehlerDifferential From 566f81bb753bb19522bea1a6b9a750cdabf731fe Mon Sep 17 00:00:00 2001 From: Hannah Scholz Date: Thu, 24 Oct 2024 16:05:11 +0000 Subject: [PATCH 380/505] feat: another composition of `ContinuousOn` and `Continuous` (#18164) --- Mathlib/Topology/ContinuousOn.lean | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/Mathlib/Topology/ContinuousOn.lean b/Mathlib/Topology/ContinuousOn.lean index 1de7949f03494..b9e11fa1067f8 100644 --- a/Mathlib/Topology/ContinuousOn.lean +++ b/Mathlib/Topology/ContinuousOn.lean @@ -861,6 +861,10 @@ theorem ContinuousOn.comp_continuous {g : β → γ} {f : α → β} {s : Set β rw [continuous_iff_continuousOn_univ] at * exact hg.comp hf fun x _ => hs x +theorem ContinuousOn.image_comp_continuous {g : β → γ} {f : α → β} {s : Set α} + (hg : ContinuousOn g (f '' s)) (hf : Continuous f) : ContinuousOn (g ∘ f) s := + hg.comp hf.continuousOn (s.mapsTo_image f) + theorem ContinuousAt.comp₂_continuousWithinAt {f : β × γ → δ} {g : α → β} {h : α → γ} {x : α} {s : Set α} (hf : ContinuousAt f (g x, h x)) (hg : ContinuousWithinAt g s x) (hh : ContinuousWithinAt h s x) : From 491b1755846663bf0e4ea785bcbf9a5c168e0686 Mon Sep 17 00:00:00 2001 From: Dagur Asgeirsson Date: Thu, 24 Oct 2024 16:05:12 +0000 Subject: [PATCH 381/505] chore(Condensed): resolve a few TODOs (#18167) These are not relevant now that #14027 has been merged. --- Mathlib/Condensed/Discrete/Basic.lean | 4 ++++ Mathlib/Condensed/Functors.lean | 9 --------- Mathlib/Condensed/Light/Functors.lean | 6 ------ 3 files changed, 4 insertions(+), 15 deletions(-) diff --git a/Mathlib/Condensed/Discrete/Basic.lean b/Mathlib/Condensed/Discrete/Basic.lean index d7964352317fa..72b29ed2500c3 100644 --- a/Mathlib/Condensed/Discrete/Basic.lean +++ b/Mathlib/Condensed/Discrete/Basic.lean @@ -19,6 +19,10 @@ In `Condensed.discreteUnderlyingAdj` we prove that this functor is left adjoint functor from `Condensed C` to `C`. We also give the variant `LightCondensed.discreteUnderlyingAdj` for light condensed objects. + +The file `Mathlib.Condensed.Discrete.Characterization` defines a predicate `IsDiscrete` on +condensed and and light condensed objects, and provides several conditions on a (light) condensed +set or module that characterize it as discrete. -/ universe u v w diff --git a/Mathlib/Condensed/Functors.lean b/Mathlib/Condensed/Functors.lean index c83e320059c0d..0fe59dc5e8866 100644 --- a/Mathlib/Condensed/Functors.lean +++ b/Mathlib/Condensed/Functors.lean @@ -20,15 +20,6 @@ sets. * `compHausToCondensed : CompHaus.{u} ⥤ CondensedSet.{u}` is essentially the yoneda presheaf functor. We also define `profiniteToCondensed` and `stoneanToCondensed`. -TODO (Dagur): - -* Define the analogues of `compHausToCondensed` for sheaves on `Profinite` and `Stonean` and provide - the relevant isomorphisms with `profiniteToCondensed` and `stoneanToCondensed`. - -* Define the functor `Type (u+1) ⥤ CondensedSet.{u}` which takes a set `X` to the presheaf given by - mapping a compact Hausdorff space `S` to `LocallyConstant S X`, along with the isomorphism with - the functor that goes through `TopCat.{u+1}`. - -/ universe u v diff --git a/Mathlib/Condensed/Light/Functors.lean b/Mathlib/Condensed/Light/Functors.lean index 7d58dbe10bb93..f8996768aa631 100644 --- a/Mathlib/Condensed/Light/Functors.lean +++ b/Mathlib/Condensed/Light/Functors.lean @@ -17,12 +17,6 @@ sets. * `lightProfiniteToLightCondSet : LightProfinite.{u} ⥤ LightCondSet.{u}` is the yoneda presheaf functor. -TODO (Dagur): - -* Define the functor `Type u ⥤ LightCondSet.{u}` which takes a set `X` to the presheaf given by - mapping a light profinite space `S` to `LocallyConstant S X`, along with the isomorphism with - the functor that goes through `TopCat.{u+1}`. - -/ universe u v From 8786716ec7f5a6ceb3f131f69c29213a71071515 Mon Sep 17 00:00:00 2001 From: Peter Nelson <71660771+apnelson1@users.noreply.github.com> Date: Thu, 24 Oct 2024 16:05:14 +0000 Subject: [PATCH 382/505] chore(Data/Matroid/Basic): document suffix convention (#18177) We add a few lines to the docstring for `Matroid` indicating the fact that theorem names use matroid predicates on sets as suffixes rather than prefixes. --- Mathlib/Data/Matroid/Basic.lean | 6 +++++- 1 file changed, 5 insertions(+), 1 deletion(-) diff --git a/Mathlib/Data/Matroid/Basic.lean b/Mathlib/Data/Matroid/Basic.lean index 69571ed264c5e..9031fedef544f 100644 --- a/Mathlib/Data/Matroid/Basic.lean +++ b/Mathlib/Data/Matroid/Basic.lean @@ -133,7 +133,7 @@ There are a few design decisions worth discussing. so our notation mirrors common practice. ### Notation - We use a couple of nonstandard conventions in theorem names that are related to the above. + We use a few nonstandard conventions in theorem names that are related to the above. First, we mirror common informal practice by referring explicitly to the `ground` set rather than the notation `E`. (Writing `ground` everywhere in a proof term would be unwieldy, and writing `E` in theorem names would be unnatural to read.) @@ -146,6 +146,10 @@ There are a few design decisions worth discussing. with respect to the ground set, rather than a complement within a type. The lemma `compl_base_dual` is one of the many examples of this. + Finally, in theorem names, matroid predicates that apply to sets + (such as `Base`, `Indep`, `Basis`) are typically used as suffixes rather than prefixes. + For instance, we have `ground_indep_iff_base` rather than `indep_ground_iff_base`. + ## References [1] The standard text on matroid theory From 6937f0825aff0dbc61a002a817fb41c2745b0720 Mon Sep 17 00:00:00 2001 From: Markus Himmel Date: Thu, 24 Oct 2024 16:27:26 +0000 Subject: [PATCH 383/505] feat: Ind-objects are closed under filtered colimits (#17451) Co-authored-by: Markus Himmel --- Mathlib.lean | 2 + .../Comma/Presheaf/Colimit.lean | 59 +++++++ .../Limits/Indization/FilteredColimits.lean | 163 ++++++++++++++++++ .../Limits/Preserves/Yoneda.lean | 2 +- 4 files changed, 225 insertions(+), 1 deletion(-) create mode 100644 Mathlib/CategoryTheory/Comma/Presheaf/Colimit.lean create mode 100644 Mathlib/CategoryTheory/Limits/Indization/FilteredColimits.lean diff --git a/Mathlib.lean b/Mathlib.lean index 2ee5eb975f465..64c30bfb19da4 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1532,6 +1532,7 @@ import Mathlib.CategoryTheory.Comma.Arrow import Mathlib.CategoryTheory.Comma.Basic import Mathlib.CategoryTheory.Comma.Over import Mathlib.CategoryTheory.Comma.Presheaf.Basic +import Mathlib.CategoryTheory.Comma.Presheaf.Colimit import Mathlib.CategoryTheory.Comma.StructuredArrow import Mathlib.CategoryTheory.ComposableArrows import Mathlib.CategoryTheory.ConcreteCategory.Basic @@ -1679,6 +1680,7 @@ import Mathlib.CategoryTheory.Limits.FunctorCategory.Finite import Mathlib.CategoryTheory.Limits.FunctorToTypes import Mathlib.CategoryTheory.Limits.HasLimits import Mathlib.CategoryTheory.Limits.IndYoneda +import Mathlib.CategoryTheory.Limits.Indization.FilteredColimits import Mathlib.CategoryTheory.Limits.Indization.IndObject import Mathlib.CategoryTheory.Limits.IsConnected import Mathlib.CategoryTheory.Limits.IsLimit diff --git a/Mathlib/CategoryTheory/Comma/Presheaf/Colimit.lean b/Mathlib/CategoryTheory/Comma/Presheaf/Colimit.lean new file mode 100644 index 0000000000000..11c42fe3b5cb7 --- /dev/null +++ b/Mathlib/CategoryTheory/Comma/Presheaf/Colimit.lean @@ -0,0 +1,59 @@ +/- +Copyright (c) 2024 Markus Himmel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Markus Himmel +-/ +import Mathlib.CategoryTheory.Comma.Presheaf.Basic +import Mathlib.CategoryTheory.Limits.Preserves.Yoneda +import Mathlib.CategoryTheory.Limits.Over + +/-! +# Relative Yoneda preserves certain colimits + +In this file we turn the statement `yonedaYonedaColimit` from +`CategoryTheory.Limits.Preserves.Yoneda` from a functor `F : J ⥤ Cᵒᵖ ⥤ Type v` into a statement +about families of presheaves over `A`, i.e., functors `F : J ⥤ Over A`. +-/ + +namespace CategoryTheory + +open Category Opposite Limits + +universe w v u + +variable {C : Type u} [Category.{v} C] {A : Cᵒᵖ ⥤ Type v} + + +variable {J : Type v} [SmallCategory J] {A : Cᵒᵖ ⥤ Type v} (F : J ⥤ Over A) + +-- We introduce some local notation to reduce visual noise in the following proof +local notation "E" => Equivalence.functor (overEquivPresheafCostructuredArrow A) +local notation "E.obj" => + Prefunctor.obj (Functor.toPrefunctor (Equivalence.functor (overEquivPresheafCostructuredArrow A))) + +/-- Naturally in `X`, we have `Hom(YX, colim_i Fi) ≅ colim_i Hom(YX, Fi)`, where `Y` is the + "Yoneda embedding" `CostructuredArrow.toOver yoneda A`. This is a relative version of + `yonedaYonedaColimit`. -/ +noncomputable def CostructuredArrow.toOverCompYonedaColimit : + (CostructuredArrow.toOver yoneda A).op ⋙ yoneda.obj (colimit F) ≅ + (CostructuredArrow.toOver yoneda A).op ⋙ colimit (F ⋙ yoneda) := calc + (CostructuredArrow.toOver yoneda A).op ⋙ yoneda.obj (colimit F) + ≅ yoneda.op ⋙ yoneda.obj (E.obj (colimit F)) := + CostructuredArrow.toOverCompYoneda A _ + _ ≅ yoneda.op ⋙ yoneda.obj (colimit (F ⋙ E)) := + isoWhiskerLeft yoneda.op (yoneda.mapIso (preservesColimitIso _ F)) + _ ≅ yoneda.op ⋙ colimit ((F ⋙ E) ⋙ yoneda) := + yonedaYonedaColimit _ + _ ≅ yoneda.op ⋙ ((F ⋙ E) ⋙ yoneda).flip ⋙ colim := + isoWhiskerLeft _ (colimitIsoFlipCompColim _) + _ ≅ (yoneda.op ⋙ coyoneda ⋙ (whiskeringLeft _ _ _).obj E) ⋙ + (whiskeringLeft _ _ _).obj F ⋙ colim := + Iso.refl _ + _ ≅ (CostructuredArrow.toOver yoneda A).op ⋙ coyoneda ⋙ (whiskeringLeft _ _ _).obj F ⋙ colim := + isoWhiskerRight (CostructuredArrow.toOverCompCoyoneda _).symm _ + _ ≅ (CostructuredArrow.toOver yoneda A).op ⋙ (F ⋙ yoneda).flip ⋙ colim := + Iso.refl _ + _ ≅ (CostructuredArrow.toOver yoneda A).op ⋙ colimit (F ⋙ yoneda) := + isoWhiskerLeft _ (colimitIsoFlipCompColim _).symm + +end CategoryTheory diff --git a/Mathlib/CategoryTheory/Limits/Indization/FilteredColimits.lean b/Mathlib/CategoryTheory/Limits/Indization/FilteredColimits.lean new file mode 100644 index 0000000000000..84bc8acad8947 --- /dev/null +++ b/Mathlib/CategoryTheory/Limits/Indization/FilteredColimits.lean @@ -0,0 +1,163 @@ +/- +Copyright (c) 2024 Markus Himmel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Markus Himmel +-/ +import Mathlib.CategoryTheory.Comma.Presheaf.Colimit +import Mathlib.CategoryTheory.Limits.Filtered +import Mathlib.CategoryTheory.Limits.FilteredColimitCommutesFiniteLimit +import Mathlib.CategoryTheory.Limits.FunctorToTypes +import Mathlib.CategoryTheory.Limits.Indization.IndObject +import Mathlib.Logic.Small.Set + +/-! +# Ind-objects are closed under filtered colimits + +We show that if `F : I ⥤ Cᵒᵖ ⥤ Type v` is a functor such that `I` is small and filtered and +`F.obj i` is an ind-object for all `i`, then `colimit F` is also an ind-object. + +Our proof is a slight variant of the proof given in Kashiwara-Schapira. + +## References +* [M. Kashiwara, P. Schapira, *Categories and Sheaves*][Kashiwara2006], Theorem 6.1.8 +-/ + +universe v u + +namespace CategoryTheory.Limits + +open CategoryTheory CategoryTheory.CostructuredArrow + +variable {C : Type u} [Category.{v} C] + +namespace IndizationClosedUnderFilteredColimitsAux + +variable {I : Type v} [SmallCategory I] (F : I ⥤ Cᵒᵖ ⥤ Type v) + + +section Interchange + +/-! +We start by stating the key interchange property `exists_nonempty_limit_obj_of_isColimit`. It +consists of pulling out a colimit out of a hom functor and interchanging a filtered colimit with +a finite limit. +-/ + +variable {J : Type v} [SmallCategory J] [FinCategory J] + +variable (G : J ⥤ CostructuredArrow yoneda (colimit F)) + +-- We introduce notation for the functor `J ⥤ Over (colimit F)` induced by `G`. +local notation "𝒢" => Functor.op G ⋙ Functor.op (toOver yoneda (colimit F)) + +variable {K : Type v} [SmallCategory K] (H : K ⥤ Over (colimit F)) + +/-- (implementation) Pulling out a colimit out of a hom functor is one half of the key lemma. Note + that all of the heavy lifting actually happens in `CostructuredArrow.toOverCompYonedaColimit` + and `yonedaYonedaColimit`. -/ +noncomputable def compYonedaColimitIsoColimitCompYoneda : + 𝒢 ⋙ yoneda.obj (colimit H) ≅ colimit (H ⋙ yoneda ⋙ (whiskeringLeft _ _ _).obj 𝒢) := calc + 𝒢 ⋙ yoneda.obj (colimit H) ≅ 𝒢 ⋙ colimit (H ⋙ yoneda) := + isoWhiskerLeft G.op (toOverCompYonedaColimit H) + _ ≅ 𝒢 ⋙ (H ⋙ yoneda).flip ⋙ colim := isoWhiskerLeft _ (colimitIsoFlipCompColim _) + _ ≅ (H ⋙ yoneda ⋙ (whiskeringLeft _ _ _).obj 𝒢).flip ⋙ colim := Iso.refl _ + _ ≅ colimit (H ⋙ yoneda ⋙ (whiskeringLeft _ _ _).obj 𝒢) := (colimitIsoFlipCompColim _).symm + +theorem exists_nonempty_limit_obj_of_colimit [IsFiltered K] + (h : Nonempty <| limit <| 𝒢 ⋙ yoneda.obj (colimit H)) : + ∃ k, Nonempty <| limit <| 𝒢 ⋙ yoneda.obj (H.obj k) := by + obtain ⟨t⟩ := h + let t₂ := limMap (compYonedaColimitIsoColimitCompYoneda F G H).hom t + let t₃ := (colimitLimitIso (H ⋙ yoneda ⋙ (whiskeringLeft _ _ _).obj 𝒢).flip).inv t₂ + obtain ⟨k, y, -⟩ := Types.jointly_surjective'.{v, max u v} t₃ + refine ⟨k, ⟨?_⟩⟩ + let z := (limitObjIsoLimitCompEvaluation (H ⋙ yoneda ⋙ (whiskeringLeft _ _ _).obj 𝒢).flip k).hom y + let y := flipCompEvaluation (H ⋙ yoneda ⋙ (whiskeringLeft _ _ _).obj 𝒢) k + exact (lim.mapIso y).hom z + +theorem exists_nonempty_limit_obj_of_isColimit [IsFiltered K] {c : Cocone H} (hc : IsColimit c) + (T : Over (colimit F)) (hT : c.pt ≅ T) + (h : Nonempty <| limit <| 𝒢 ⋙ yoneda.obj T) : + ∃ k, Nonempty <| limit <| 𝒢 ⋙ yoneda.obj (H.obj k) := by + refine exists_nonempty_limit_obj_of_colimit F G H ?_ + suffices T ≅ colimit H from Nonempty.map (lim.map (whiskerLeft 𝒢 (yoneda.map this.hom))) h + refine hT.symm ≪≫ IsColimit.coconePointUniqueUpToIso hc (colimit.isColimit _) + +end Interchange + +theorem isFiltered [IsFiltered I] (hF : ∀ i, IsIndObject (F.obj i)) : + IsFiltered (CostructuredArrow yoneda (colimit F)) := by + -- It suffices to show that for any functor `G : J ⥤ CostructuredArrow yoneda (colimit F)` with + -- `J` finite there is some `X` such that the set + -- `lim Hom_{CostructuredArrow yoneda (colimit F)}(G·, X)` is nonempty. + refine IsFiltered.iff_nonempty_limit.mpr (fun {J _ _} G => ?_) + + -- We begin by remarking that `lim Hom_{Over (colimit F)}(yG·, 𝟙 (colimit F))` is nonempty, + -- simply because `𝟙 (colimit F)` is the terminal object. Here `y` is the functor + -- `CostructuredArrow yoneda (colimit F) ⥤ Over (colimit F)` induced by `yoneda`. + have h₁ : Nonempty (limit (G.op ⋙ (toOver _ _).op ⋙ yoneda.obj (Over.mk (𝟙 (colimit F))))) := + ⟨Types.Limit.mk _ (fun j => Over.mkIdTerminal.from _) (by simp)⟩ + + -- `𝟙 (colimit F)` is the colimit of the diagram in `Over (colimit F)` given by the arrows of + -- the form `Fi ⟶ colimit F`. Thus, pulling the colimit out of the hom functor and commuting + -- the finite limit with the filtered colimit, we obtain + -- `lim_j Hom_{Over (colimit F)}(yGj, 𝟙 (colimit F)) ≅` + -- `colim_i lim_j Hom_{Over (colimit F)}(yGj, colimit.ι F i)`, and so we find `i` such that + -- the limit is non-empty. + obtain ⟨i, hi⟩ := exists_nonempty_limit_obj_of_isColimit F G _ + (colimit.isColimitToOver F) _ (Iso.refl _) h₁ + + -- `F.obj i` is a small filtered colimit of representables, say of the functor `H : K ⥤ C`, so + -- `𝟙 (F.obj i)` is the colimit of the arrows of the form `yHk ⟶ Fi` in `Over Fi`. + -- Then `colimit.ι F i` is the colimit of the arrows of the form + -- `H.obj F ⟶ F.obj i ⟶ colimit F` in `Over (colimit F)`. + obtain ⟨⟨P⟩⟩ := hF i + let hc : IsColimit ((Over.map (colimit.ι F i)).mapCocone P.cocone.toOver) := + isColimitOfPreserves (Over.map _) (Over.isColimitToOver P.coconeIsColimit) + + -- Again, we pull the colimit out of the hom functor and commute limit and colimit to obtain + -- `lim_j Hom_{Over (colimit F)}(yGj, colimit.ι F i) ≅` + -- `colim_k lim_j Hom_{Over (colimit F)}(yGj, yHk)`, and so we find `k` such that the limit + -- is non-empty. + obtain ⟨k, hk⟩ : ∃ k, Nonempty (limit (G.op ⋙ (toOver yoneda (colimit F)).op ⋙ + yoneda.obj ((toOver yoneda (colimit F)).obj <| + (pre P.F yoneda (colimit F)).obj <| (map (colimit.ι F i)).obj <| mk _))) := + exists_nonempty_limit_obj_of_isColimit F G _ hc _ (Iso.refl _) hi + + have htO : (toOver yoneda (colimit F)).FullyFaithful := .ofFullyFaithful _ + -- Since the inclusion `y : CostructuredArrow yoneda (colimit F) ⥤ Over (colimit F)` is fully + -- faithful, `lim_j Hom_{Over (colimit F)}(yGj, yHk) ≅` + -- `lim_j Hom_{CostructuredArrow yoneda (colimit F)}(Gj, Hk)` and so `Hk` is the object we're + -- looking for. + let q := htO.homNatIsoMaxRight + obtain ⟨t'⟩ := Nonempty.map (limMap (isoWhiskerLeft G.op (q _)).hom) hk + exact ⟨_, ⟨((preservesLimitIso uliftFunctor.{u, v} _).inv t').down⟩⟩ + +end IndizationClosedUnderFilteredColimitsAux + +theorem isIndObject_colimit (I : Type v) [SmallCategory I] [IsFiltered I] + (F : I ⥤ Cᵒᵖ ⥤ Type v) (hF : ∀ i, IsIndObject (F.obj i)) : IsIndObject (colimit F) := by + have : IsFiltered (CostructuredArrow yoneda (colimit F)) := + IndizationClosedUnderFilteredColimitsAux.isFiltered F hF + refine (isIndObject_iff _).mpr ⟨this, ?_⟩ + + -- It remains to show that `CostructuredArrow yoneda (colimit F)` is finally small. Because we + -- have already shown it is filtered, it suffices to exhibit a small weakly terminal set. For this + -- we use that all the `CostructuredArrow yoneda (F.obj i)` have small weakly terminal sets. + have : ∀ i, ∃ (s : Set (CostructuredArrow yoneda (F.obj i))) (_ : Small.{v} s), + ∀ i, ∃ j ∈ s, Nonempty (i ⟶ j) := + fun i => (hF i).finallySmall.exists_small_weakly_terminal_set + choose s hs j hjs hj using this + refine finallySmall_of_small_weakly_terminal_set + (⋃ i, (map (colimit.ι F i)).obj '' (s i)) (fun A => ?_) + obtain ⟨i, y, hy⟩ := FunctorToTypes.jointly_surjective'.{v, v} F _ (yonedaEquiv A.hom) + let y' : CostructuredArrow yoneda (F.obj i) := mk (yonedaEquiv.symm y) + obtain ⟨x⟩ := hj _ y' + refine ⟨(map (colimit.ι F i)).obj (j i y'), ?_, ⟨?_⟩⟩ + · simp only [Set.mem_iUnion, Set.mem_image] + exact ⟨i, j i y', hjs _ _, rfl⟩ + · refine ?_ ≫ (map (colimit.ι F i)).map x + refine homMk (𝟙 A.left) (yonedaEquiv.injective ?_) + simp [-EmbeddingLike.apply_eq_iff_eq, hy, yonedaEquiv_comp, y'] + +end CategoryTheory.Limits diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Yoneda.lean b/Mathlib/CategoryTheory/Limits/Preserves/Yoneda.lean index 629c22852c6a5..e0da3d65a8d90 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Yoneda.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Yoneda.lean @@ -23,7 +23,7 @@ pointwise. ## See also There is also a relative version of this statement where `F : J ⥤ Over A` for some presheaf -`A`, see `CategoryTheory.Comma.Presheaf`. +`A`, see `CategoryTheory.Comma.Presheaf.Colimit`. -/ From 201aff369a90cac045403f10f08d36a2b6c4473a Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Thu, 24 Oct 2024 16:27:28 +0000 Subject: [PATCH 384/505] feat(Algebra/CharP/Two): more lemmas on char two (#17483) Co-authored-by: Ruben Van de Velde --- Mathlib/Algebra/CharP/Two.lean | 40 ++++++++++++++++++++++++++++---- Mathlib/Algebra/Ring/Parity.lean | 32 +++++++++++++++++++++++++ 2 files changed, 67 insertions(+), 5 deletions(-) diff --git a/Mathlib/Algebra/CharP/Two.lean b/Mathlib/Algebra/CharP/Two.lean index 6a4a8f2d12bda..edd4d620d33e7 100644 --- a/Mathlib/Algebra/CharP/Two.lean +++ b/Mathlib/Algebra/CharP/Two.lean @@ -22,15 +22,34 @@ variable {R ι : Type*} namespace CharTwo +section AddMonoidWithOne + +variable [AddMonoidWithOne R] + +theorem two_eq_zero [CharP R 2] : (2 : R) = 0 := by + rw [← Nat.cast_two, CharP.cast_eq_zero] + +/-- The only hypotheses required to build a `CharP R 2` instance are `1 ≠ 0` and `2 = 0`. -/ +theorem of_one_ne_zero_of_two_eq_zero (h₁ : (1 : R) ≠ 0) (h₂ : (2 : R) = 0) : CharP R 2 where + cast_eq_zero_iff' n := by + obtain hn | hn := Nat.even_or_odd n + · simp_rw [hn.two_dvd, iff_true] + exact natCast_eq_zero_of_even_of_two_eq_zero hn h₂ + · simp_rw [hn.not_two_dvd_nat, iff_false] + rwa [natCast_eq_one_of_odd_of_two_eq_zero hn h₂] + +end AddMonoidWithOne + section Semiring variable [Semiring R] [CharP R 2] -theorem two_eq_zero : (2 : R) = 0 := by rw [← Nat.cast_two, CharP.cast_eq_zero] - @[scoped simp] theorem add_self_eq_zero (x : R) : x + x = 0 := by rw [← two_smul R x, two_eq_zero, zero_smul] +@[scoped simp] +protected theorem two_nsmul (x : R) : 2 • x = 0 := by rw [two_smul, add_self_eq_zero] + end Semiring section Ring @@ -39,7 +58,7 @@ variable [Ring R] [CharP R 2] @[scoped simp] theorem neg_eq (x : R) : -x = x := by - rw [neg_eq_iff_add_eq_zero, ← two_smul R x, two_eq_zero, zero_smul] + rw [neg_eq_iff_add_eq_zero, add_self_eq_zero] theorem neg_eq' : Neg.neg = (id : R → R) := funext neg_eq @@ -47,8 +66,19 @@ theorem neg_eq' : Neg.neg = (id : R → R) := @[scoped simp] theorem sub_eq_add (x y : R) : x - y = x + y := by rw [sub_eq_add_neg, neg_eq] -theorem sub_eq_add' : HSub.hSub = ((· + ·) : R → R → R) := - funext fun x => funext fun y => sub_eq_add x y +@[deprecated sub_eq_add (since := "2024-10-24")] +theorem sub_eq_add' : HSub.hSub = (· + · : R → R → R) := + funext₂ sub_eq_add + +theorem add_eq_iff_eq_add {a b c : R} : a + b = c ↔ a = c + b := by + rw [← sub_eq_iff_eq_add, sub_eq_add] + +theorem eq_add_iff_add_eq {a b c : R} : a = b + c ↔ a + c = b := by + rw [← eq_sub_iff_add_eq, sub_eq_add] + +@[scoped simp] +protected theorem two_zsmul (x : R) : (2 : ℤ) • x = 0 := by + rw [two_zsmul, add_self_eq_zero] end Ring diff --git a/Mathlib/Algebra/Ring/Parity.lean b/Mathlib/Algebra/Ring/Parity.lean index 6adab468d1261..b4890ab73e4bc 100644 --- a/Mathlib/Algebra/Ring/Parity.lean +++ b/Mathlib/Algebra/Ring/Parity.lean @@ -238,6 +238,9 @@ lemma even_xor_odd' (n : ℕ) : ∃ k, Xor' (n = 2 * k) (n = 2 * k + 1) := by · simpa only [← two_mul, eq_self_iff_true, xor_true] using (succ_ne_self (2 * k)).symm · simpa only [xor_true, xor_comm] using (succ_ne_self _) +lemma odd_add_one {n : ℕ} : Odd (n + 1) ↔ ¬ Odd n := by + rw [← not_even_iff_odd, Nat.even_add_one, not_even_iff_odd] + 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 rw [odd_iff.1 hm, even_iff.1 (hm.add_odd hn)] @@ -361,3 +364,32 @@ lemma neg_one_pow_eq_one_iff_even {R : Type*} [Monoid R] [HasDistribNeg R] {n : (h : (-1 : R) ≠ 1) : (-1 : R) ^ n = 1 ↔ Even n where mp h' := of_not_not fun hn ↦ h <| (not_even_iff_odd.1 hn).neg_one_pow.symm.trans h' mpr := Even.neg_one_pow + +section CharTwo + +-- We state the following theorems in terms of the slightly more general `2 = 0` hypothesis. + +variable {R : Type*} [AddMonoidWithOne R] + +private theorem natCast_eq_zero_or_one_of_two_eq_zero' (n : ℕ) (h : (2 : R) = 0) : + (Even n → (n : R) = 0) ∧ (Odd n → (n : R) = 1) := by + induction n using Nat.twoStepInduction with + | zero => simp + | one => simp + | more n _ _ => simpa [add_assoc, Nat.even_add_one, Nat.odd_add_one, h] + +theorem natCast_eq_zero_of_even_of_two_eq_zero {n : ℕ} (hn : Even n) (h : (2 : R) = 0) : + (n : R) = 0 := + (natCast_eq_zero_or_one_of_two_eq_zero' n h).1 hn + +theorem natCast_eq_one_of_odd_of_two_eq_zero {n : ℕ} (hn : Odd n) (h : (2 : R) = 0) : + (n : R) = 1 := + (natCast_eq_zero_or_one_of_two_eq_zero' n h).2 hn + +theorem natCast_eq_zero_or_one_of_two_eq_zero (n : ℕ) (h : (2 : R) = 0) : + (n : R) = 0 ∨ (n : R) = 1 := by + obtain hn | hn := Nat.even_or_odd n + · exact Or.inl <| natCast_eq_zero_of_even_of_two_eq_zero hn h + · exact Or.inr <| natCast_eq_one_of_odd_of_two_eq_zero hn h + +end CharTwo From 62a047f83bd145703c99ab38fbd336960f2239bf Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Thu, 24 Oct 2024 16:27:29 +0000 Subject: [PATCH 385/505] feat: SMul instance for AddSubmonoid and lemmas (#18096) + Introduce a `SMul` instance for two `AddSubmonoid`s and provide lemmas + Introduce a `Mul` instance for `AddSubgroup` + Prove lemmas for the existing `Mul` instance on `AddSubmonoid`, which are generalized from corresponding `Submodule` lemmas --- Mathlib/Algebra/Group/Subgroup/Pointwise.lean | 20 ++++++ .../Algebra/Group/Submonoid/Pointwise.lean | 70 +++++++++++++++++-- 2 files changed, 85 insertions(+), 5 deletions(-) diff --git a/Mathlib/Algebra/Group/Subgroup/Pointwise.lean b/Mathlib/Algebra/Group/Subgroup/Pointwise.lean index f42601fa4cc09..0378c1dd57b4b 100644 --- a/Mathlib/Algebra/Group/Subgroup/Pointwise.lean +++ b/Mathlib/Algebra/Group/Subgroup/Pointwise.lean @@ -547,4 +547,24 @@ theorem le_pointwise_smul_iff₀ {a : α} (ha : a ≠ 0) {S T : AddSubgroup A} : end GroupWithZero +section Mul + +variable {R : Type*} [NonUnitalNonAssocRing R] + +/-- For additive subgroups `S` and `T` of a ring, the product of `S` and `T` as submonoids +is automatically a subgroup, which we define as the product of `S` and `T` as subgroups. -/ +protected def mul : Mul (AddSubgroup R) where + mul M N := + { __ := M.toAddSubmonoid * N.toAddSubmonoid + neg_mem' := fun h ↦ AddSubmonoid.mul_induction_on h + (fun m hm n hn ↦ by rw [← neg_mul]; exact AddSubmonoid.mul_mem_mul (M.neg_mem hm) hn) + fun r₁ r₂ h₁ h₂ ↦ by rw [neg_add]; exact (M.1 * N.1).add_mem h₁ h₂ } + +scoped[Pointwise] attribute [instance] AddSubgroup.mul + +theorem mul_toAddSubmonoid (M N : AddSubgroup R) : + (M * N).toAddSubmonoid = M.toAddSubmonoid * N.toAddSubmonoid := rfl + +end Mul + end AddSubgroup diff --git a/Mathlib/Algebra/Group/Submonoid/Pointwise.lean b/Mathlib/Algebra/Group/Submonoid/Pointwise.lean index fea277c83d7d4..74f18f06313d8 100644 --- a/Mathlib/Algebra/Group/Submonoid/Pointwise.lean +++ b/Mathlib/Algebra/Group/Submonoid/Pointwise.lean @@ -488,20 +488,47 @@ theorem bot_mul (S : AddSubmonoid R) : ⊥ * S = ⊥ := eq_bot_iff.2 <| mul_le.2 fun m hm n _ => by rw [AddSubmonoid.mem_bot] at hm ⊢; rw [hm, zero_mul] +variable {M N P Q : AddSubmonoid R} + @[mono] -theorem mul_le_mul {M N P Q : AddSubmonoid R} (hmp : M ≤ P) (hnq : N ≤ Q) : M * N ≤ P * Q := +theorem mul_le_mul (hmp : M ≤ P) (hnq : N ≤ Q) : M * N ≤ P * Q := mul_le.2 fun _m hm _n hn => mul_mem_mul (hmp hm) (hnq hn) -theorem mul_le_mul_left {M N P : AddSubmonoid R} (h : M ≤ N) : M * P ≤ N * P := +theorem mul_le_mul_left (h : M ≤ N) : M * P ≤ N * P := mul_le_mul h (le_refl P) -theorem mul_le_mul_right {M N P : AddSubmonoid R} (h : N ≤ P) : M * N ≤ M * P := +theorem mul_le_mul_right (h : N ≤ P) : M * N ≤ M * P := mul_le_mul (le_refl M) h -theorem mul_subset_mul {M N : AddSubmonoid R} : - (↑M : Set R) * (↑N : Set R) ⊆ (↑(M * N) : Set R) := +theorem mul_subset_mul : (↑M : Set R) * (↑N : Set R) ⊆ (↑(M * N) : Set R) := mul_subset_iff.2 fun _i hi _j hj ↦ mul_mem_mul hi hj +theorem mul_sup : M * (N ⊔ P) = M * N ⊔ M * P := + le_antisymm (mul_le.mpr fun m hm np hnp ↦ by + obtain ⟨n, hn, p, hp, rfl⟩ := mem_sup.mp hnp + rw [left_distrib]; exact add_mem_sup (mul_mem_mul hm hn) <| mul_mem_mul hm hp) + (sup_le (mul_le_mul_right le_sup_left) <| mul_le_mul_right le_sup_right) + +theorem sup_mul : (M ⊔ N) * P = M * P ⊔ N * P := + le_antisymm (mul_le.mpr fun mn hmn p hp ↦ by + obtain ⟨m, hm, n, hn, rfl⟩ := mem_sup.mp hmn + rw [right_distrib]; exact add_mem_sup (mul_mem_mul hm hp) <| mul_mem_mul hn hp) + (sup_le (mul_le_mul_left le_sup_left) <| mul_le_mul_left le_sup_right) + +variable {ι : Sort*} + +theorem iSup_mul (S : ι → AddSubmonoid R) (T : AddSubmonoid R) : (⨆ i, S i) * T = ⨆ i, S i * T := + le_antisymm (mul_le.mpr fun s hs t ht ↦ iSup_induction _ (C := (· * t ∈ _)) hs + (fun i s hs ↦ mem_iSup_of_mem i <| mul_mem_mul hs ht) (by simp_rw [zero_mul]; apply zero_mem) + fun _ _ ↦ by simp_rw [right_distrib]; apply add_mem) <| + iSup_le fun i ↦ mul_le_mul_left (le_iSup _ i) + +theorem mul_iSup (T : AddSubmonoid R) (S : ι → AddSubmonoid R) : (T * ⨆ i, S i) = ⨆ i, T * S i := + le_antisymm (mul_le.mpr fun t ht s hs ↦ iSup_induction _ (C := (t * · ∈ _)) hs + (fun i s hs ↦ mem_iSup_of_mem i <| mul_mem_mul ht hs) (by simp_rw [mul_zero]; apply zero_mem) + fun _ _ ↦ by simp_rw [left_distrib]; apply add_mem) <| + iSup_le fun i ↦ mul_le_mul_right (le_iSup _ i) + end NonUnitalNonAssocSemiring section NonUnitalNonAssocRing @@ -589,6 +616,39 @@ theorem pow_subset_pow {s : AddSubmonoid R} {n : ℕ} : (↑s : Set R) ^ n ⊆ end Semiring +section SMul + +variable [AddMonoid R] [DistribSMul R A] + +/-- For `M : Submonoid R` and `N : AddSubmonoid A`, `M • N` is the additive submonoid +generated by all `m • n` where `m ∈ M` and `n ∈ N`. -/ +protected def smul : SMul (AddSubmonoid R) (AddSubmonoid A) where + smul M N := ⨆ s : M, N.map (DistribSMul.toAddMonoidHom A s.1) + +scoped[Pointwise] attribute [instance] AddSubmonoid.smul + +example {R} [Semiring R] : Mul.toSMul (AddSubmonoid R) = AddSubmonoid.smul := rfl + +variable {M M' : AddSubmonoid R} {N P : AddSubmonoid A} {m : R} {n : A} + +theorem smul_mem_smul (hm : m ∈ M) (hn : n ∈ N) : m • n ∈ M • N := + (le_iSup _ ⟨m, hm⟩ : _ ≤ M • N) ⟨n, hn, by rfl⟩ + +theorem smul_le : M • N ≤ P ↔ ∀ m ∈ M, ∀ n ∈ N, m • n ∈ P := + ⟨fun H _m hm _n hn => H <| smul_mem_smul hm hn, fun H => + iSup_le fun ⟨m, hm⟩ => map_le_iff_le_comap.2 fun n hn => H m hm n hn⟩ + +@[elab_as_elim] +protected theorem smul_induction_on {C : A → Prop} {a : A} (ha : a ∈ M • N) + (hm : ∀ m ∈ M, ∀ n ∈ N, C (m • n)) (hadd : ∀ x y, C x → C y → C (x + y)) : C a := + (@smul_le _ _ _ _ _ _ _ ⟨⟨setOf C, hadd _ _⟩, by + simpa only [smul_zero] using hm _ (zero_mem _) _ (zero_mem _)⟩).2 hm ha + +theorem smul_le_smul (h : M ≤ M') (hnp : N ≤ P) : M • N ≤ M' • P := + smul_le.2 fun _m hm _n hn => smul_mem_smul (h hm) (hnp hn) + +end SMul + end AddSubmonoid namespace Set.IsPWO From 9715344187cb8e034f5192efd474d00faee444f0 Mon Sep 17 00:00:00 2001 From: Riccardo Brasca Date: Thu, 24 Oct 2024 16:57:55 +0000 Subject: [PATCH 386/505] feat: miscellaneous results about finrank (#18064) From `flt-regular`. --- .../LinearAlgebra/Dimension/RankNullity.lean | 52 +++++++++++++++---- Mathlib/LinearAlgebra/FiniteDimensional.lean | 8 --- Mathlib/NumberTheory/RamificationInertia.lean | 2 +- 3 files changed, 44 insertions(+), 18 deletions(-) diff --git a/Mathlib/LinearAlgebra/Dimension/RankNullity.lean b/Mathlib/LinearAlgebra/Dimension/RankNullity.lean index 2e6cb6ba25096..79f9779cf6271 100644 --- a/Mathlib/LinearAlgebra/Dimension/RankNullity.lean +++ b/Mathlib/LinearAlgebra/Dimension/RankNullity.lean @@ -24,7 +24,7 @@ See `nonempty_oreSet_of_strongRankCondition` for a start. -/ universe u v -open Function Set Cardinal +open Function Set Cardinal Submodule LinearMap variable {R} {M M₁ M₂ M₃ : Type u} {M' : Type v} [Ring R] variable [AddCommGroup M] [AddCommGroup M₁] [AddCommGroup M₂] [AddCommGroup M₃] [AddCommGroup M'] @@ -49,7 +49,7 @@ class HasRankNullity (R : Type v) [inst : Ring R] : Prop where variable [HasRankNullity.{u} R] -lemma rank_quotient_add_rank (N : Submodule R M) : +lemma Submodule.rank_quotient_add_rank (N : Submodule R M) : Module.rank R (M ⧸ N) + Module.rank R N = Module.rank R M := HasRankNullity.rank_quotient_add_rank N @@ -66,24 +66,24 @@ theorem nontrivial_of_hasRankNullity : Nontrivial R := by attribute [local instance] nontrivial_of_hasRankNullity -theorem lift_rank_range_add_rank_ker (f : M →ₗ[R] M') : +theorem LinearMap.lift_rank_range_add_rank_ker (f : M →ₗ[R] M') : lift.{u} (Module.rank R (LinearMap.range f)) + lift.{v} (Module.rank R (LinearMap.ker f)) = lift.{v} (Module.rank R M) := by haveI := fun p : Submodule R M => Classical.decEq (M ⧸ p) rw [← f.quotKerEquivRange.lift_rank_eq, ← lift_add, rank_quotient_add_rank] /-- The **rank-nullity theorem** -/ -theorem rank_range_add_rank_ker (f : M →ₗ[R] M₁) : +theorem LinearMap.rank_range_add_rank_ker (f : M →ₗ[R] M₁) : Module.rank R (LinearMap.range f) + Module.rank R (LinearMap.ker f) = Module.rank R M := by haveI := fun p : Submodule R M => Classical.decEq (M ⧸ p) rw [← f.quotKerEquivRange.rank_eq, rank_quotient_add_rank] -theorem lift_rank_eq_of_surjective {f : M →ₗ[R] M'} (h : Surjective f) : +theorem LinearMap.lift_rank_eq_of_surjective {f : M →ₗ[R] M'} (h : Surjective f) : lift.{v} (Module.rank R M) = lift.{u} (Module.rank R M') + lift.{v} (Module.rank R (LinearMap.ker f)) := by rw [← lift_rank_range_add_rank_ker f, ← rank_range_of_surjective f h] -theorem rank_eq_of_surjective {f : M →ₗ[R] M₁} (h : Surjective f) : +theorem LinearMap.rank_eq_of_surjective {f : M →ₗ[R] M₁} (h : Surjective f) : Module.rank R M = Module.rank R M₁ + Module.rank R (LinearMap.ker f) := by rw [← rank_range_add_rank_ker f, ← rank_range_of_surjective f h] @@ -92,7 +92,7 @@ theorem exists_linearIndependent_of_lt_rank [StrongRankCondition R] ∃ t, s ⊆ t ∧ #t = Module.rank R M ∧ LinearIndependent (ι := t) R Subtype.val := by obtain ⟨t, ht, ht'⟩ := exists_set_linearIndependent R (M ⧸ Submodule.span R s) choose sec hsec using Submodule.Quotient.mk_surjective (Submodule.span R s) - have hsec' : Submodule.Quotient.mk ∘ sec = id := funext hsec + have hsec' : Submodule.Quotient.mk ∘ sec = _root_.id := funext hsec have hst : Disjoint s (sec '' t) := by rw [Set.disjoint_iff] rintro _ ⟨hxs, ⟨x, hxt, rfl⟩⟩ @@ -140,8 +140,8 @@ theorem exists_linearIndependent_pair_of_one_lt_rank [StrongRankCondition R] rw [this] at hy exact ⟨y, hy⟩ -theorem exists_smul_not_mem_of_rank_lt {N : Submodule R M} (h : Module.rank R N < Module.rank R M) : - ∃ m : M, ∀ r : R, r ≠ 0 → r • m ∉ N := by +theorem Submodule.exists_smul_not_mem_of_rank_lt {N : Submodule R M} + (h : Module.rank R N < Module.rank R M) : ∃ m : M, ∀ r : R, r ≠ 0 → r • m ∉ N := by have : Module.rank R (M ⧸ N) ≠ 0 := by intro e rw [← rank_quotient_add_rank N, e, zero_add] at h @@ -196,4 +196,38 @@ theorem exists_linearIndependent_pair_of_one_lt_finrank [NoZeroSMulDivisors R M] ∃ y, LinearIndependent R ![x, y] := exists_linearIndependent_pair_of_one_lt_rank (one_lt_rank_of_one_lt_finrank h) hx +/-- Rank-nullity theorem using `finrank`. -/ +lemma Submodule.finrank_quotient_add_finrank [Module.Finite R M] (N : Submodule R M) : + finrank R (M ⧸ N) + finrank R N = finrank R M := by + rw [← Nat.cast_inj (R := Cardinal), Module.finrank_eq_rank, Nat.cast_add, Module.finrank_eq_rank, + Submodule.finrank_eq_rank] + exact HasRankNullity.rank_quotient_add_rank _ + + +/-- Rank-nullity theorem using `finrank` and subtraction. -/ +lemma Submodule.finrank_quotient [Module.Finite R M] {S : Type*} [Ring S] [SMul R S] [Module S M] + [IsScalarTower R S M] (N : Submodule S M) : finrank R (M ⧸ N) = finrank R M - finrank R N := by + rw [← (N.restrictScalars R).finrank_quotient_add_finrank] + exact Nat.eq_sub_of_add_eq rfl + end Finrank + +section + +open Submodule Module + +variable [StrongRankCondition R] [Module.Finite R M] + +lemma Submodule.exists_of_finrank_lt (N : Submodule R M) (h : finrank R N < finrank R M) : + ∃ m : M, ∀ r : R, r ≠ 0 → r • m ∉ N := by + obtain ⟨s, hs, hs'⟩ := + exists_finset_linearIndependent_of_le_finrank (R := R) (M := M ⧸ N) le_rfl + obtain ⟨v, hv⟩ : s.Nonempty := by rwa [Finset.nonempty_iff_ne_empty, ne_eq, ← Finset.card_eq_zero, + hs, finrank_quotient, tsub_eq_zero_iff_le, not_le] + obtain ⟨v, rfl⟩ := N.mkQ_surjective v + refine ⟨v, fun r hr ↦ mt ?_ hr⟩ + have := linearIndependent_iff.mp hs' (Finsupp.single ⟨_, hv⟩ r) + rwa [Finsupp.linearCombination_single, Finsupp.single_eq_zero, ← LinearMap.map_smul, + Submodule.mkQ_apply, Submodule.Quotient.mk_eq_zero] at this + +end diff --git a/Mathlib/LinearAlgebra/FiniteDimensional.lean b/Mathlib/LinearAlgebra/FiniteDimensional.lean index 7b2f68f6cdcdd..3b375180e50b1 100644 --- a/Mathlib/LinearAlgebra/FiniteDimensional.lean +++ b/Mathlib/LinearAlgebra/FiniteDimensional.lean @@ -33,14 +33,6 @@ section DivisionRing variable [DivisionRing K] [AddCommGroup V] [Module K V] -/-- In a finite-dimensional vector space, the dimensions of a submodule and of the corresponding -quotient add up to the dimension of the space. -/ -theorem finrank_quotient_add_finrank [FiniteDimensional K V] (s : Submodule K V) : - finrank K (V ⧸ s) + finrank K s = finrank K V := by - have := rank_quotient_add_rank s - rw [← finrank_eq_rank, ← finrank_eq_rank, ← finrank_eq_rank] at this - exact mod_cast this - /-- The dimension of a strict submodule is strictly bounded by the dimension of the ambient space. -/ theorem finrank_lt [FiniteDimensional K V] {s : Submodule K V} (h : s < ⊤) : diff --git a/Mathlib/NumberTheory/RamificationInertia.lean b/Mathlib/NumberTheory/RamificationInertia.lean index a69dfce45bd5a..71fc4378b8d5f 100644 --- a/Mathlib/NumberTheory/RamificationInertia.lean +++ b/Mathlib/NumberTheory/RamificationInertia.lean @@ -588,7 +588,7 @@ theorem rank_pow_quot_aux [IsDedekindDomain S] [p.IsMaximal] [P.IsPrime] (hP0 : letI : Field (R ⧸ p) := Ideal.Quotient.field _ rw [← rank_range_of_injective _ (powQuotSuccInclusion_injective f p P i), (quotientRangePowQuotSuccInclusionEquiv f p P hP0 hi).symm.rank_eq] - exact (rank_quotient_add_rank (LinearMap.range (powQuotSuccInclusion f p P i))).symm + exact (Submodule.rank_quotient_add_rank (LinearMap.range (powQuotSuccInclusion f p P i))).symm theorem rank_pow_quot [IsDedekindDomain S] [p.IsMaximal] [P.IsPrime] (hP0 : P ≠ ⊥) (i : ℕ) (hi : i ≤ e) : From 8fd75accaff95050c8b594847f537bdde54a54ef Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Thu, 24 Oct 2024 18:27:41 +0000 Subject: [PATCH 387/505] chore: clean up "dangerous instance" porting notes (#18041) Most of these notes are about issues in Lean 3 that are no longer issues in Lean 4, so they can be easily removed. A few need some more explanation. --- Mathlib/Algebra/Algebra/Equiv.lean | 3 --- Mathlib/Algebra/Algebra/Hom.lean | 3 --- Mathlib/Algebra/Algebra/NonUnitalHom.lean | 4 ---- Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean | 1 - Mathlib/Algebra/Module/LinearMap/Defs.lean | 5 ----- Mathlib/Algebra/Ring/Hom/Defs.lean | 2 +- Mathlib/Analysis/Seminorm.lean | 3 --- Mathlib/CategoryTheory/Comma/Arrow.lean | 1 - Mathlib/GroupTheory/GroupAction/Hom.lean | 6 ------ Mathlib/LinearAlgebra/Ray.lean | 2 +- Mathlib/Order/Heyting/Hom.lean | 4 ++-- 11 files changed, 4 insertions(+), 30 deletions(-) diff --git a/Mathlib/Algebra/Algebra/Equiv.lean b/Mathlib/Algebra/Algebra/Equiv.lean index f5e674a2f3566..6e995f594999b 100644 --- a/Mathlib/Algebra/Algebra/Equiv.lean +++ b/Mathlib/Algebra/Algebra/Equiv.lean @@ -44,8 +44,6 @@ class AlgEquivClass (F : Type*) (R A B : outParam Type*) [CommSemiring R] [Semir /-- An equivalence of algebras commutes with the action of scalars. -/ commutes : ∀ (f : F) (r : R), f (algebraMap R A r) = algebraMap R B r --- Porting note: Removed nolint dangerousInstance from AlgEquivClass.toRingEquivClass - namespace AlgEquivClass -- See note [lower instance priority] @@ -141,7 +139,6 @@ protected theorem coe_coe {F : Type*} [EquivLike F A₁ A₂] [AlgEquivClass F R theorem coe_fun_injective : @Function.Injective (A₁ ≃ₐ[R] A₂) (A₁ → A₂) fun e => (e : A₁ → A₂) := DFunLike.coe_injective --- Porting note: Made to CoeOut instance from Coe, not dangerous anymore instance hasCoeToRingEquiv : CoeOut (A₁ ≃ₐ[R] A₂) (A₁ ≃+* A₂) := ⟨AlgEquiv.toRingEquiv⟩ diff --git a/Mathlib/Algebra/Algebra/Hom.lean b/Mathlib/Algebra/Algebra/Hom.lean index b9f3122d4c8bd..0f6f35fd279e4 100644 --- a/Mathlib/Algebra/Algebra/Hom.lean +++ b/Mathlib/Algebra/Algebra/Hom.lean @@ -45,9 +45,6 @@ class AlgHomClass (F : Type*) (R A B : outParam Type*) [FunLike F A B] extends RingHomClass F A B : Prop where commutes : ∀ (f : F) (r : R), f (algebraMap R A r) = algebraMap R B r --- Porting note: `dangerousInstance` linter has become smarter about `outParam`s --- attribute [nolint dangerousInstance] AlgHomClass.toRingHomClass - -- For now, don't replace `AlgHom.commutes` and `AlgHomClass.commutes` with the more generic lemma. -- The file `Mathlib.NumberTheory.NumberField.CanonicalEmbedding.FundamentalCone` slows down by -- 15% if we would do so (see benchmark on PR #18040). diff --git a/Mathlib/Algebra/Algebra/NonUnitalHom.lean b/Mathlib/Algebra/Algebra/NonUnitalHom.lean index f91df6f1aa69a..678960a93a6bf 100644 --- a/Mathlib/Algebra/Algebra/NonUnitalHom.lean +++ b/Mathlib/Algebra/Algebra/NonUnitalHom.lean @@ -80,12 +80,8 @@ abbrev NonUnitalAlgHomClass (F : Type*) (R A B : outParam Type*) [DistribMulAction R A] [DistribMulAction R B] [FunLike F A B] := NonUnitalAlgSemiHomClass F (MonoidHom.id R) A B --- Porting note: commented out, not dangerous --- attribute [nolint dangerousInstance] NonUnitalAlgHomClass.toMulHomClass - namespace NonUnitalAlgHomClass --- Porting note: Made following instance non-dangerous through [...] -> [...] replacement -- See note [lower instance priority] instance (priority := 100) toNonUnitalRingHomClass {F R S A B : Type*} {_ : Monoid R} {_ : Monoid S} {φ : outParam (R →* S)} diff --git a/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean b/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean index 6e0cef1234889..5b6bf266bb5de 100644 --- a/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean +++ b/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean @@ -1190,7 +1190,6 @@ end Group section VSub --- Porting note: Reordered [VSub α β] and [DecidableEq α] to make vsub less dangerous. Bad? variable [VSub α β] [DecidableEq α] {s s₁ s₂ t t₁ t₂ : Finset β} {u : Finset α} {a : α} {b c : β} /-- The pointwise subtraction of two finsets `s` and `t`: `s -ᵥ t = {x -ᵥ y | x ∈ s, y ∈ t}`. -/ diff --git a/Mathlib/Algebra/Module/LinearMap/Defs.lean b/Mathlib/Algebra/Module/LinearMap/Defs.lean index 8e3800a3eed03..f0138c3409d11 100644 --- a/Mathlib/Algebra/Module/LinearMap/Defs.lean +++ b/Mathlib/Algebra/Module/LinearMap/Defs.lean @@ -110,10 +110,6 @@ class SemilinearMapClass (F : Type*) {R S : outParam Type*} [Semiring R] [Semiri end --- Porting note: `dangerousInstance` linter has become smarter about `outParam`s --- `σ` becomes a metavariable but that's fine because it's an `outParam` --- attribute [nolint dangerousInstance] SemilinearMapClass.toAddHomClass - -- `map_smulₛₗ` should be `@[simp]` but doesn't fire due to `lean4#3701`. -- attribute [simp] map_smulₛₗ @@ -139,7 +135,6 @@ variable [AddCommMonoid M] [AddCommMonoid M₁] [AddCommMonoid M₂] [AddCommMon variable [Module R M] [Module R M₂] [Module S M₃] variable {σ : R →+* S} --- Porting note: the `dangerousInstance` linter has become smarter about `outParam`s instance (priority := 100) instAddMonoidHomClass [FunLike F M M₃] [SemilinearMapClass F σ M M₃] : AddMonoidHomClass F M M₃ := { SemilinearMapClass.toAddHomClass with diff --git a/Mathlib/Algebra/Ring/Hom/Defs.lean b/Mathlib/Algebra/Ring/Hom/Defs.lean index 2ba354beb18de..c03030622109a 100644 --- a/Mathlib/Algebra/Ring/Hom/Defs.lean +++ b/Mathlib/Algebra/Ring/Hom/Defs.lean @@ -329,7 +329,7 @@ class RingHomClass (F : Type*) (α β : outParam Type*) variable [FunLike F α β] --- Porting note: marked `{}` rather than `[]` to prevent dangerous instances +-- See note [implicit instance arguments]. variable {_ : NonAssocSemiring α} {_ : NonAssocSemiring β} [RingHomClass F α β] /-- Turn an element of a type `F` satisfying `RingHomClass F α β` into an actual diff --git a/Mathlib/Analysis/Seminorm.lean b/Mathlib/Analysis/Seminorm.lean index 751860a0f37b2..1c096d7eb792e 100644 --- a/Mathlib/Analysis/Seminorm.lean +++ b/Mathlib/Analysis/Seminorm.lean @@ -61,9 +61,6 @@ class SeminormClass (F : Type*) (𝕜 E : outParam Type*) [SeminormedRing 𝕜] export SeminormClass (map_smul_eq_mul) --- Porting note: dangerous instances no longer exist --- attribute [nolint dangerousInstance] SeminormClass.toAddGroupSeminormClass - section Of /-- Alternative constructor for a `Seminorm` on an `AddCommGroup E` that is a module over a diff --git a/Mathlib/CategoryTheory/Comma/Arrow.lean b/Mathlib/CategoryTheory/Comma/Arrow.lean index 2d8f03930b499..4abf32df55770 100644 --- a/Mathlib/CategoryTheory/Comma/Arrow.lean +++ b/Mathlib/CategoryTheory/Comma/Arrow.lean @@ -88,7 +88,6 @@ theorem mk_injective (A B : T) : theorem mk_inj (A B : T) {f g : A ⟶ B} : Arrow.mk f = Arrow.mk g ↔ f = g := (mk_injective A B).eq_iff -/- Porting note: was marked as dangerous instance so changed from `Coe` to `CoeOut` -/ instance {X Y : T} : CoeOut (X ⟶ Y) (Arrow T) where coe := mk diff --git a/Mathlib/GroupTheory/GroupAction/Hom.lean b/Mathlib/GroupTheory/GroupAction/Hom.lean index fdebbaf6ec0f9..962885a423eb3 100644 --- a/Mathlib/GroupTheory/GroupAction/Hom.lean +++ b/Mathlib/GroupTheory/GroupAction/Hom.lean @@ -356,9 +356,6 @@ abbrev DistribMulActionHomClass (F : Type*) (M : outParam Type*) [DistribMulAction M A] [DistribMulAction M B] [FunLike F A B] := DistribMulActionSemiHomClass F (MonoidHom.id M) A B -/- porting note: Removed a @[nolint dangerousInstance] for -DistribMulActionHomClass.toAddMonoidHomClass not dangerous due to `outParam`s -/ - namespace DistribMulActionHom /- Porting note (#11215): TODO decide whether the next two instances should be removed @@ -601,9 +598,6 @@ abbrev MulSemiringActionHomClass [DistribMulAction M R] [DistribMulAction M S] [FunLike F R S] := MulSemiringActionSemiHomClass F (MonoidHom.id M) R S -/- porting note: Removed a @[nolint dangerousInstance] for MulSemiringActionHomClass.toRingHomClass - not dangerous due to outParam -/ - namespace MulSemiringActionHom /- Porting note (#11215): TODO decide whether the next two instances should be removed diff --git a/Mathlib/LinearAlgebra/Ray.lean b/Mathlib/LinearAlgebra/Ray.lean index 680c02e7722d3..9fdfae1150fba 100644 --- a/Mathlib/LinearAlgebra/Ray.lean +++ b/Mathlib/LinearAlgebra/Ray.lean @@ -188,9 +188,9 @@ set_option linter.unusedVariables false in def RayVector (R M : Type*) [Zero M] := { v : M // v ≠ 0 } --- Porting note: Made Coe into CoeOut so it's not dangerous anymore instance RayVector.coe [Zero M] : CoeOut (RayVector R M) M where coe := Subtype.val + instance {R M : Type*} [Zero M] [Nontrivial M] : Nonempty (RayVector R M) := let ⟨x, hx⟩ := exists_ne (0 : M) ⟨⟨x, hx⟩⟩ diff --git a/Mathlib/Order/Heyting/Hom.lean b/Mathlib/Order/Heyting/Hom.lean index ab4677d3ab3cb..6a1e4f16688ce 100644 --- a/Mathlib/Order/Heyting/Hom.lean +++ b/Mathlib/Order/Heyting/Hom.lean @@ -99,8 +99,8 @@ section Hom variable [FunLike F α β] -/- Porting note: `[HeytingAlgebra α, β]` -> `{ _ : HeytingAlgebra α, β}` as a dangerous instance fix -similar for Coheyting & Biheyting instances -/ +/-! This section passes in some instances implicitly. See note [implicit instance arguments] -/ + -- See note [lower instance priority] instance (priority := 100) HeytingHomClass.toBoundedLatticeHomClass [HeytingAlgebra α] { _ : HeytingAlgebra β} [HeytingHomClass F α β] : BoundedLatticeHomClass F α β := From b8fa953fd3a165a7adf4d7d94ff265bcf093c739 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Bj=C3=B8rn=20Kjos-Hanssen?= Date: Thu, 24 Oct 2024 19:02:22 +0000 Subject: [PATCH 388/505] feat: Equivalence between OnePoint K and projective line over field K (#17266) Adding equivalence between OnePoint K and projective line over K, as discussed at https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/.E2.9C.94.20Projective.20extended.20real.20line/near/463620378 The [first PR for this](https://github.com/leanprover-community/mathlib4/pull/16152) became unable to automatically merge, so I am closing it and opening up this one. (I do have a proof extending this to a homeomorphism for K = Real, but perhaps it is more disciplined to do the Equiv first?) Co-authored-by: Oliver Nash <7734364+ocfnash@users.noreply.github.com> Co-authored-by: Riccardo Brasca --- Mathlib.lean | 1 + .../LinearAlgebra/Projectivization/Basic.lean | 14 ++++ .../Compactification/OnePointEquiv.lean | 72 +++++++++++++++++++ 3 files changed, 87 insertions(+) create mode 100644 Mathlib/Topology/Compactification/OnePointEquiv.lean diff --git a/Mathlib.lean b/Mathlib.lean index 64c30bfb19da4..5c1737f14e072 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4706,6 +4706,7 @@ import Mathlib.Topology.Clopen import Mathlib.Topology.ClopenBox import Mathlib.Topology.CompactOpen import Mathlib.Topology.Compactification.OnePoint +import Mathlib.Topology.Compactification.OnePointEquiv import Mathlib.Topology.Compactness.Compact import Mathlib.Topology.Compactness.CompactlyGeneratedSpace import Mathlib.Topology.Compactness.Exterior diff --git a/Mathlib/LinearAlgebra/Projectivization/Basic.lean b/Mathlib/LinearAlgebra/Projectivization/Basic.lean index 94d4e7925bb43..9a12ad4ccb657 100644 --- a/Mathlib/LinearAlgebra/Projectivization/Basic.lean +++ b/Mathlib/LinearAlgebra/Projectivization/Basic.lean @@ -68,6 +68,20 @@ instance [Nontrivial V] : Nonempty (ℙ K V) := variable {K} +/-- A function on non-zero vectors which is independent of scale, descends to a function on the +projectivization. -/ +protected def lift {α : Type*} (f : { v : V // v ≠ 0 } → α) + (hf : ∀ (a b : { v : V // v ≠ 0 }) (t : K), a = t • (b : V) → f a = f b) + (x : ℙ K V) : α := + Quotient.lift f (by rintro ⟨-, hv⟩ ⟨w, hw⟩ ⟨⟨t, -⟩, rfl⟩; exact hf ⟨_, hv⟩ ⟨w, hw⟩ t rfl) x + +@[simp] +protected lemma lift_mk {α : Type*} (f : { v : V // v ≠ 0 } → α) + (hf : ∀ (a b : { v : V // v ≠ 0 }) (t : K), a = t • (b : V) → f a = f b) + (v : V) (hv : v ≠ 0) : + Projectivization.lift f hf (mk K v hv) = f ⟨v, hv⟩ := + rfl + /-- Choose a representative of `v : Projectivization K V` in `V`. -/ protected noncomputable def rep (v : ℙ K V) : V := v.out' diff --git a/Mathlib/Topology/Compactification/OnePointEquiv.lean b/Mathlib/Topology/Compactification/OnePointEquiv.lean new file mode 100644 index 0000000000000..4282e6d8b60e4 --- /dev/null +++ b/Mathlib/Topology/Compactification/OnePointEquiv.lean @@ -0,0 +1,72 @@ +/- +Copyright (c) 2024 Bjørn Kjos-Hanssen. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Bjørn Kjos-Hanssen, Oliver Nash +-/ +import Mathlib.LinearAlgebra.Projectivization.Basic +import Mathlib.Topology.Compactification.OnePoint + +/-! +# One-point compactification and projectivization + +We construct a set-theoretic equivalence between +`OnePoint K` and the projectivization `ℙ K (K × K)` for an arbitrary division ring `K`. + +TODO: Add the extension of this equivalence to a homeomorphism in the case `K = ℝ`, +where `OnePoint ℝ` gets the topology of one-point compactification. + + +## Main definitions and results + +* `OnePoint.equivProjectivization` : the equivalence `OnePoint K ≃ ℙ K (K × K)`. + +## Tags + +one-point extension, projectivization +-/ + +open scoped LinearAlgebra.Projectivization OnePoint +open Projectivization + +namespace OnePoint +variable (K : Type*) [DivisionRing K] [DecidableEq K] + +/-- The one-point compactification of a division ring `K` is equivalent to + the projectivivization `ℙ K (K × K)`. -/ +def equivProjectivization : + OnePoint K ≃ ℙ K (K × K) where + toFun p := Option.elim p (mk K (1, 0) (by simp)) (fun t ↦ mk K (t, 1) (by simp)) + invFun p := by + refine Projectivization.lift + (fun u : {v : K × K // v ≠ 0} ↦ if u.1.2 ≠ 0 then ((u.1.2)⁻¹ * u.1.1) else ∞) ?_ p + rintro ⟨-, hv⟩ ⟨⟨x, y⟩, hw⟩ t rfl + have ht : t ≠ 0 := by rintro rfl; simp at hv + by_cases h₀ : y = 0 <;> simp [h₀, ht, mul_assoc] + left_inv p := by cases p <;> simp [OnePoint.infty, OnePoint.some] + right_inv p := by + induction' p using ind with p hp + obtain ⟨x, y⟩ := p + by_cases h₀ : y = 0 <;> + simp only [Option.elim, ne_eq, ite_not, h₀, Projectivization.lift_mk, reduceIte] + · have h₀' : x ≠ 0 := by aesop + simp only [mk_eq_mk_iff, Prod.smul_mk, smul_zero, Prod.mk.injEq, and_true] + exact ⟨Units.mk0 _ (inv_ne_zero h₀'), by simp [h₀']⟩ + · simp only [mk_eq_mk_iff, Prod.smul_mk, Prod.mk.injEq] + exact ⟨Units.mk0 _ (inv_ne_zero h₀), by simp [h₀]⟩ + +@[simp] +lemma equivProjectivization_apply_infinity : + equivProjectivization K ∞ = mk K ⟨1, 0⟩ (by simp) := + rfl + +@[simp] +lemma equivProjectivization_apply_coe (t : K) : + equivProjectivization K t = mk K ⟨t, 1⟩ (by simp) := + rfl + +@[simp] +lemma equivProjectivization_symm_apply_mk (x y : K) (h : (x, y) ≠ 0) : + (equivProjectivization K).symm (mk K ⟨x, y⟩ h) = if y = 0 then ∞ else y⁻¹ * x := by + simp [equivProjectivization] + +end OnePoint From 2473c8327c4c5123084773ae280bae5b7f0954a8 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Thu, 24 Oct 2024 19:02:24 +0000 Subject: [PATCH 389/505] feat(SetTheory/Cardinal/Basic): there exists a well-founded linear ordering of any type (#18183) --- Mathlib/SetTheory/Cardinal/Basic.lean | 6 ++++++ Mathlib/Topology/EMetricSpace/Paracompact.lean | 10 ++++------ 2 files changed, 10 insertions(+), 6 deletions(-) diff --git a/Mathlib/SetTheory/Cardinal/Basic.lean b/Mathlib/SetTheory/Cardinal/Basic.lean index cfd91460ac93e..c5202a59be16d 100644 --- a/Mathlib/SetTheory/Cardinal/Basic.lean +++ b/Mathlib/SetTheory/Cardinal/Basic.lean @@ -971,6 +971,12 @@ instance WellOrderingRel.isWellOrder : IsWellOrder α WellOrderingRel := instance IsWellOrder.subtype_nonempty : Nonempty { r // IsWellOrder α r } := ⟨⟨WellOrderingRel, inferInstance⟩⟩ +variable (α) in +/-- The **well-ordering theorem** (or **Zermelo's theorem**): every type has a well-order -/ +theorem exists_wellOrder : ∃ (_ : LinearOrder α), WellFoundedLT α := by + classical + exact ⟨linearOrderOfSTO WellOrderingRel, WellOrderingRel.isWellOrder.toIsWellFounded⟩ + /-! ### Small sets of cardinals -/ namespace Cardinal diff --git a/Mathlib/Topology/EMetricSpace/Paracompact.lean b/Mathlib/Topology/EMetricSpace/Paracompact.lean index f7e7082bc036a..b725de087262c 100644 --- a/Mathlib/Topology/EMetricSpace/Paracompact.lean +++ b/Mathlib/Topology/EMetricSpace/Paracompact.lean @@ -48,14 +48,12 @@ instance (priority := 100) instParacompactSpace [PseudoEMetricSpace α] : Paraco refine ⟨fun ι s ho hcov => ?_⟩ simp only [iUnion_eq_univ_iff] at hcov -- choose a well founded order on `S` - -- Porting note (#11215): TODO: add lemma that claims `∃ i : LinearOrder ι, WellFoundedLT ι` - let _ : LinearOrder ι := by classical exact linearOrderOfSTO WellOrderingRel - have wf : WellFounded ((· < ·) : ι → ι → Prop) := @IsWellFounded.wf ι WellOrderingRel _ + obtain ⟨_, wf⟩ := exists_wellOrder ι -- Let `ind x` be the minimal index `s : S` such that `x ∈ s`. - set ind : α → ι := fun x => wf.min { i : ι | x ∈ s i } (hcov x) - have mem_ind : ∀ x, x ∈ s (ind x) := fun x => wf.min_mem _ (hcov x) + set ind : α → ι := fun x => wellFounded_lt.min { i : ι | x ∈ s i } (hcov x) + have mem_ind : ∀ x, x ∈ s (ind x) := fun x => wellFounded_lt.min_mem _ (hcov x) have nmem_of_lt_ind : ∀ {x i}, i < ind x → x ∉ s i := @fun x i hlt hxi => - wf.not_lt_min _ (hcov x) hxi hlt + wellFounded_lt.not_lt_min _ (hcov x) hxi hlt /- The refinement `D : ℕ → ι → Set α` is defined recursively. For each `n` and `i`, `D n i` is the union of balls `ball x (1 / 2 ^ n)` over all points `x` such that From 280b21599476b03427c097e46da069277be22558 Mon Sep 17 00:00:00 2001 From: damiano Date: Thu, 24 Oct 2024 19:02:25 +0000 Subject: [PATCH 390/505] chore: remove a variable and collapse rest (#18194) This is just the removal of the variable `D`, followed by the renaming of `E --> D`, `F --> E`, `G --> F`, all happening automatically. #17715 --- .../Calculus/ContDiff/FiniteDimension.lean | 23 ++++++++++--------- 1 file changed, 12 insertions(+), 11 deletions(-) diff --git a/Mathlib/Analysis/Calculus/ContDiff/FiniteDimension.lean b/Mathlib/Analysis/Calculus/ContDiff/FiniteDimension.lean index 5b3ed3d9a9e4f..e5d673c7ffac0 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/FiniteDimension.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/FiniteDimension.lean @@ -14,11 +14,12 @@ import Mathlib.Analysis.Normed.Module.FiniteDimension noncomputable section -universe uD uE uF uG +universe uD uE uF -variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {D : Type uD} [NormedAddCommGroup D] - [NormedSpace 𝕜 D] {E : Type uE} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {F : Type uF} - [NormedAddCommGroup F] [NormedSpace 𝕜 F] {G : Type uG} [NormedAddCommGroup G] [NormedSpace 𝕜 G] +variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] + {D : Type uD} [NormedAddCommGroup D] [NormedSpace 𝕜 D] + {E : Type uE} [NormedAddCommGroup E] [NormedSpace 𝕜 E] + {F : Type uF} [NormedAddCommGroup F] [NormedSpace 𝕜 F] /-! ### Finite dimensional results -/ @@ -29,17 +30,17 @@ open Function Module variable [CompleteSpace 𝕜] /-- A family of continuous linear maps is `C^n` on `s` if all its applications are. -/ -theorem contDiffOn_clm_apply {n : ℕ∞} {f : E → F →L[𝕜] G} {s : Set E} [FiniteDimensional 𝕜 F] : +theorem contDiffOn_clm_apply {n : ℕ∞} {f : D → E →L[𝕜] F} {s : Set D} [FiniteDimensional 𝕜 E] : ContDiffOn 𝕜 n f s ↔ ∀ y, ContDiffOn 𝕜 n (fun x => f x y) s := by refine ⟨fun h y => h.clm_apply contDiffOn_const, fun h => ?_⟩ - let d := finrank 𝕜 F + let d := finrank 𝕜 E have hd : d = finrank 𝕜 (Fin d → 𝕜) := (finrank_fin_fun 𝕜).symm let e₁ := ContinuousLinearEquiv.ofFinrankEq hd - let e₂ := (e₁.arrowCongr (1 : G ≃L[𝕜] G)).trans (ContinuousLinearEquiv.piRing (Fin d)) + let e₂ := (e₁.arrowCongr (1 : F ≃L[𝕜] F)).trans (ContinuousLinearEquiv.piRing (Fin d)) rw [← id_comp f, ← e₂.symm_comp_self] exact e₂.symm.contDiff.comp_contDiffOn (contDiffOn_pi.mpr fun i => h _) -theorem contDiff_clm_apply_iff {n : ℕ∞} {f : E → F →L[𝕜] G} [FiniteDimensional 𝕜 F] : +theorem contDiff_clm_apply_iff {n : ℕ∞} {f : D → E →L[𝕜] F} [FiniteDimensional 𝕜 E] : ContDiff 𝕜 n f ↔ ∀ y, ContDiff 𝕜 n fun x => f x y := by simp_rw [← contDiffOn_univ, contDiffOn_clm_apply] @@ -52,16 +53,16 @@ often requires an inconvenient need to generalize `F`, which results in universe (see the discussion in the section of `ContDiff.comp`). This lemma avoids these universe issues, but only applies for finite dimensional `E`. -/ -theorem contDiff_succ_iff_fderiv_apply [FiniteDimensional 𝕜 E] {n : ℕ} {f : E → F} : +theorem contDiff_succ_iff_fderiv_apply [FiniteDimensional 𝕜 D] {n : ℕ} {f : D → E} : ContDiff 𝕜 (n + 1 : ℕ) f ↔ Differentiable 𝕜 f ∧ ∀ y, ContDiff 𝕜 n fun x => fderiv 𝕜 f x y := by rw [contDiff_succ_iff_fderiv, contDiff_clm_apply_iff] -theorem contDiffOn_succ_of_fderiv_apply [FiniteDimensional 𝕜 E] {n : ℕ} {f : E → F} {s : Set E} +theorem contDiffOn_succ_of_fderiv_apply [FiniteDimensional 𝕜 D] {n : ℕ} {f : D → E} {s : Set D} (hf : DifferentiableOn 𝕜 f s) (h : ∀ y, ContDiffOn 𝕜 n (fun x => fderivWithin 𝕜 f s x y) s) : ContDiffOn 𝕜 (n + 1 : ℕ) f s := contDiffOn_succ_of_fderivWithin hf <| contDiffOn_clm_apply.mpr h -theorem contDiffOn_succ_iff_fderiv_apply [FiniteDimensional 𝕜 E] {n : ℕ} {f : E → F} {s : Set E} +theorem contDiffOn_succ_iff_fderiv_apply [FiniteDimensional 𝕜 D] {n : ℕ} {f : D → E} {s : Set D} (hs : UniqueDiffOn 𝕜 s) : ContDiffOn 𝕜 (n + 1 : ℕ) f s ↔ DifferentiableOn 𝕜 f s ∧ ∀ y, ContDiffOn 𝕜 n (fun x => fderivWithin 𝕜 f s x y) s := by From 779f4ef188b4f5e6e4774515005b39787ee6725e Mon Sep 17 00:00:00 2001 From: sgouezel Date: Thu, 24 Oct 2024 19:28:36 +0000 Subject: [PATCH 391/505] chore: make the model with corners implicit in differential geometry statements (#17687) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Currently, the model with corners `I` is explicit in many (most?) differential geometry statements. However, when rewriting, or applying theorems, it can be inferred from the context 95% of the time. Therefore, we make it implicit in most statements: the rule of thumb is that in `variable` lines we should have `{I : ModelWithCorners 𝕜 E H}`, and use `variable (I) in` for definitions. This looks like a clear improvement to me, making proofs more readable and avoiding clutter. For the few cases where we need to provide it explicitly, named arguments work well. No mathematical change at all, just arguments implicitness. --- .../Complex/UpperHalfPlane/Manifold.lean | 2 +- .../Geometry/Manifold/AnalyticManifold.lean | 8 +- Mathlib/Geometry/Manifold/BumpFunction.lean | 20 +- Mathlib/Geometry/Manifold/Complex.lean | 6 +- .../Geometry/Manifold/ContMDiff/Atlas.lean | 21 +- .../Geometry/Manifold/ContMDiff/Basic.lean | 24 +- Mathlib/Geometry/Manifold/ContMDiff/Defs.lean | 97 +++++---- .../Geometry/Manifold/ContMDiff/Product.lean | 9 +- .../Geometry/Manifold/ContMDiffMFDeriv.lean | 34 ++- Mathlib/Geometry/Manifold/ContMDiffMap.lean | 8 +- Mathlib/Geometry/Manifold/Diffeomorph.lean | 4 +- Mathlib/Geometry/Manifold/Instances/Real.lean | 4 +- .../Instances/UnitsOfNormedAlgebra.lean | 4 +- Mathlib/Geometry/Manifold/IntegralCurve.lean | 26 +-- .../Geometry/Manifold/InteriorBoundary.lean | 30 +-- Mathlib/Geometry/Manifold/MFDeriv/Atlas.lean | 33 +-- Mathlib/Geometry/Manifold/MFDeriv/Basic.lean | 24 +- Mathlib/Geometry/Manifold/MFDeriv/Defs.lean | 24 +- .../Manifold/MFDeriv/SpecificFunctions.lean | 94 ++++---- .../Geometry/Manifold/MFDeriv/Tangent.lean | 8 +- .../Manifold/MFDeriv/UniqueDifferential.lean | 10 +- .../Geometry/Manifold/PartitionOfUnity.lean | 5 +- Mathlib/Geometry/Manifold/Sheaf/Smooth.lean | 10 +- .../Manifold/SmoothManifoldWithCorners.lean | 205 +++++++++--------- .../Geometry/Manifold/VectorBundle/Basic.lean | 54 +++-- .../Geometry/Manifold/VectorBundle/Hom.lean | 10 +- .../Manifold/VectorBundle/Pullback.lean | 4 +- .../Manifold/VectorBundle/Tangent.lean | 57 +++-- .../Geometry/Manifold/WhitneyEmbedding.lean | 4 +- Mathlib/NumberTheory/ModularForms/Basic.lean | 6 +- 30 files changed, 419 insertions(+), 426 deletions(-) diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/Manifold.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/Manifold.lean index 45772d8d62ff6..aede75f666ab9 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/Manifold.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/Manifold.lean @@ -22,7 +22,7 @@ noncomputable instance : ChartedSpace ℂ ℍ := UpperHalfPlane.isOpenEmbedding_coe.singletonChartedSpace instance : SmoothManifoldWithCorners 𝓘(ℂ) ℍ := - UpperHalfPlane.isOpenEmbedding_coe.singleton_smoothManifoldWithCorners 𝓘(ℂ) + UpperHalfPlane.isOpenEmbedding_coe.singleton_smoothManifoldWithCorners /-- The inclusion map `ℍ → ℂ` is a smooth map of manifolds. -/ theorem smooth_coe : Smooth 𝓘(ℂ) 𝓘(ℂ) ((↑) : ℍ → ℂ) := fun _ => contMDiffAt_extChartAt diff --git a/Mathlib/Geometry/Manifold/AnalyticManifold.lean b/Mathlib/Geometry/Manifold/AnalyticManifold.lean index 8ef606a4540a1..9a1ab30e1daa9 100644 --- a/Mathlib/Geometry/Manifold/AnalyticManifold.lean +++ b/Mathlib/Geometry/Manifold/AnalyticManifold.lean @@ -29,7 +29,7 @@ open scoped Manifold Filter Topology variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] variable {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {H : Type*} - [TopologicalSpace H] (I : ModelWithCorners 𝕜 E H) {M : Type*} [TopologicalSpace M] + [TopologicalSpace H] {I : ModelWithCorners 𝕜 E H} {M : Type*} [TopologicalSpace M] /-! ## `analyticGroupoid` @@ -41,6 +41,7 @@ analytic on the interior, and map the interior to itself. This allows us to def section analyticGroupoid +variable (I) in /-- Given a model with corners `(E, H)`, we define the pregroupoid of analytic transformations of `H` as the maps that are `AnalyticOn` when read in `E` through `I`. Using `AnalyticOn` rather than `AnalyticOnNhd` gives us meaningful definitions at boundary points. -/ @@ -74,6 +75,7 @@ def analyticPregroupoid : Pregroupoid H where simp only [mfld_simps, ← hx] at hy1 ⊢ rw [fg _ hy1] +variable (I) in /-- Given a model with corners `(E, H)`, we define the groupoid of analytic transformations of `H` as the maps that are `AnalyticOn` when read in `E` through `I`. Using `AnalyticOn` rather than `AnalyticOnNhd` gives us meaningful definitions at boundary points. -/ @@ -95,7 +97,7 @@ theorem symm_trans_mem_analyticGroupoid (e : PartialHomeomorph M H) : e.symm.trans e ∈ analyticGroupoid I := haveI : e.symm.trans e ≈ PartialHomeomorph.ofSet e.target e.open_target := PartialHomeomorph.symm_trans_self _ - StructureGroupoid.mem_of_eqOnSource _ (ofSet_mem_analyticGroupoid I e.open_target) this + StructureGroupoid.mem_of_eqOnSource _ (ofSet_mem_analyticGroupoid e.open_target) this /-- The analytic groupoid is closed under restriction. -/ instance : ClosedUnderRestriction (analyticGroupoid I) := @@ -103,7 +105,7 @@ instance : ClosedUnderRestriction (analyticGroupoid I) := (by rw [StructureGroupoid.le_iff] rintro e ⟨s, hs, hes⟩ - exact (analyticGroupoid I).mem_of_eqOnSource' _ _ (ofSet_mem_analyticGroupoid I hs) hes) + exact (analyticGroupoid I).mem_of_eqOnSource' _ _ (ofSet_mem_analyticGroupoid hs) hes) /-- `f ∈ analyticGroupoid` iff it and its inverse are analytic within `range I`. -/ lemma mem_analyticGroupoid {I : ModelWithCorners 𝕜 E H} {f : PartialHomeomorph H H} : diff --git a/Mathlib/Geometry/Manifold/BumpFunction.lean b/Mathlib/Geometry/Manifold/BumpFunction.lean index 16f7d52a36180..fa6698b53773a 100644 --- a/Mathlib/Geometry/Manifold/BumpFunction.lean +++ b/Mathlib/Geometry/Manifold/BumpFunction.lean @@ -95,7 +95,7 @@ theorem ball_inter_range_eq_ball_inter_target : ball (extChartAt I c c) f.rOut ∩ range I = ball (extChartAt I c c) f.rOut ∩ (extChartAt I c).target := (subset_inter inter_subset_left f.ball_subset).antisymm <| inter_subset_inter_right _ <| - extChartAt_target_subset_range _ _ + extChartAt_target_subset_range _ section FiniteDimensional @@ -120,7 +120,7 @@ theorem support_eq_inter_preimage : theorem isOpen_support : IsOpen (support f) := by rw [support_eq_inter_preimage] - exact isOpen_extChartAt_preimage I c isOpen_ball + exact isOpen_extChartAt_preimage c isOpen_ball theorem support_eq_symm_image : support f = (extChartAt I c).symm '' (ball (extChartAt I c c) f.rOut ∩ range I) := by @@ -157,7 +157,7 @@ theorem le_one : f x ≤ 1 := theorem eventuallyEq_one_of_dist_lt (hs : x ∈ (chartAt H c).source) (hd : dist (extChartAt I c x) (extChartAt I c c) < f.rIn) : f =ᶠ[𝓝 x] 1 := by - filter_upwards [IsOpen.mem_nhds (isOpen_extChartAt_preimage I c isOpen_ball) ⟨hs, hd⟩] + filter_upwards [IsOpen.mem_nhds (isOpen_extChartAt_preimage c isOpen_ball) ⟨hs, hd⟩] rintro z ⟨hzs, hzd⟩ exact f.one_of_dist_le hzs <| le_of_lt hzd @@ -183,7 +183,7 @@ theorem nonempty_support : (support f).Nonempty := theorem isCompact_symm_image_closedBall : IsCompact ((extChartAt I c).symm '' (closedBall (extChartAt I c c) f.rOut ∩ range I)) := ((isCompact_closedBall _ _).inter_right I.isClosed_range).image_of_continuousOn <| - (continuousOn_extChartAt_symm _ _).mono f.closedBall_subset + (continuousOn_extChartAt_symm _).mono f.closedBall_subset end FiniteDimensional @@ -194,7 +194,7 @@ theorem nhdsWithin_range_basis : (𝓝[range I] extChartAt I c c).HasBasis (fun _ : SmoothBumpFunction I c => True) fun f => closedBall (extChartAt I c c) f.rOut ∩ range I := by refine ((nhdsWithin_hasBasis nhds_basis_closedBall _).restrict_subset - (extChartAt_target_mem_nhdsWithin _ _)).to_hasBasis' ?_ ?_ + (extChartAt_target_mem_nhdsWithin _)).to_hasBasis' ?_ ?_ · rintro R ⟨hR0, hsub⟩ exact ⟨⟨⟨R / 2, R, half_pos hR0, half_lt_self hR0⟩, hsub⟩, trivial, Subset.rfl⟩ · exact fun f _ => inter_mem (mem_nhdsWithin_of_mem_nhds <| closedBall_mem_nhds _ f.rOut_pos) @@ -206,7 +206,7 @@ theorem isClosed_image_of_isClosed {s : Set M} (hsc : IsClosed s) (hs : s ⊆ su IsClosed (extChartAt I c '' s) := by rw [f.image_eq_inter_preimage_of_subset_support hs] refine ContinuousOn.preimage_isClosed_of_isClosed - ((continuousOn_extChartAt_symm _ _).mono f.closedBall_subset) ?_ hsc + ((continuousOn_extChartAt_symm _).mono f.closedBall_subset) ?_ hsc exact IsClosed.inter isClosed_ball I.isClosed_range /-- If `f` is a smooth bump function and `s` closed subset of the support of `f` (i.e., of the open @@ -260,8 +260,6 @@ protected theorem hasCompactSupport : HasCompactSupport f := f.isCompact_symm_image_closedBall.of_isClosed_subset isClosed_closure f.tsupport_subset_symm_image_closedBall -variable (I) - variable (c) in /-- The closures of supports of smooth bump functions centered at `c` form a basis of `𝓝 c`. In other words, each of these closures is a neighborhood of `c` and each neighborhood of `c` @@ -271,7 +269,7 @@ theorem nhds_basis_tsupport : have : (𝓝 c).HasBasis (fun _ : SmoothBumpFunction I c => True) fun f => (extChartAt I c).symm '' (closedBall (extChartAt I c c) f.rOut ∩ range I) := by - rw [← map_extChartAt_symm_nhdsWithin_range I c] + rw [← map_extChartAt_symm_nhdsWithin_range (I := I) c] exact nhdsWithin_range_basis.map _ exact this.to_hasBasis' (fun f _ => ⟨f, trivial, f.tsupport_subset_symm_image_closedBall⟩) fun f _ => f.tsupport_mem_nhds @@ -282,10 +280,10 @@ neighborhood of `c` and each neighborhood of `c` includes `support f` for some `f : SmoothBumpFunction I c` such that `tsupport f ⊆ s`. -/ theorem nhds_basis_support {s : Set M} (hs : s ∈ 𝓝 c) : (𝓝 c).HasBasis (fun f : SmoothBumpFunction I c => tsupport f ⊆ s) fun f => support f := - ((nhds_basis_tsupport I c).restrict_subset hs).to_hasBasis' + ((nhds_basis_tsupport c).restrict_subset hs).to_hasBasis' (fun f hf => ⟨f, hf.2, subset_closure⟩) fun f _ => f.support_mem_nhds -variable [SmoothManifoldWithCorners I M] {I} +variable [SmoothManifoldWithCorners I M] /-- A smooth bump function is infinitely smooth. -/ protected theorem smooth : Smooth I 𝓘(ℝ) f := by diff --git a/Mathlib/Geometry/Manifold/Complex.lean b/Mathlib/Geometry/Manifold/Complex.lean index a9ea94cdde69c..4d7fa4e80bc64 100644 --- a/Mathlib/Geometry/Manifold/Complex.lean +++ b/Mathlib/Geometry/Manifold/Complex.lean @@ -55,10 +55,10 @@ theorem Complex.norm_eventually_eq_of_mdifferentiableAt_of_isLocalMax {f : M → have hI : range I = univ := ModelWithCorners.Boundaryless.range_eq_univ have H₁ : 𝓝[range I] (e c) = 𝓝 (e c) := by rw [hI, nhdsWithin_univ] have H₂ : map e.symm (𝓝 (e c)) = 𝓝 c := by - rw [← map_extChartAt_symm_nhdsWithin_range I c, H₁] + rw [← map_extChartAt_symm_nhdsWithin_range (I := I) c, H₁] rw [← H₂, eventually_map] replace hd : ∀ᶠ y in 𝓝 (e c), DifferentiableAt ℂ (f ∘ e.symm) y := by - have : e.target ∈ 𝓝 (e c) := H₁ ▸ extChartAt_target_mem_nhdsWithin I c + have : e.target ∈ 𝓝 (e c) := H₁ ▸ extChartAt_target_mem_nhdsWithin c filter_upwards [this, Tendsto.eventually H₂.le hd] with y hyt hy₂ have hys : e.symm y ∈ (chartAt H c).source := by rw [← extChartAt_source I c] @@ -68,7 +68,7 @@ theorem Complex.norm_eventually_eq_of_mdifferentiableAt_of_isLocalMax {f : M → e.right_inv hyt] at hy₂ exact hy₂.2 convert norm_eventually_eq_of_isLocalMax hd _ - · exact congr_arg f (extChartAt_to_inv _ _).symm + · exact congr_arg f (extChartAt_to_inv _).symm · simpa only [e, IsLocalMax, IsMaxFilter, ← H₂, (· ∘ ·), extChartAt_to_inv] using hc /-! diff --git a/Mathlib/Geometry/Manifold/ContMDiff/Atlas.lean b/Mathlib/Geometry/Manifold/ContMDiff/Atlas.lean index 61057854d30c6..c4d0007f3c70a 100644 --- a/Mathlib/Geometry/Manifold/ContMDiff/Atlas.lean +++ b/Mathlib/Geometry/Manifold/ContMDiff/Atlas.lean @@ -35,7 +35,7 @@ section Atlas theorem contMDiff_model : ContMDiff I 𝓘(𝕜, E) n I := by intro x - refine (contMDiffAt_iff _ _).mpr ⟨I.continuousAt, ?_⟩ + refine contMDiffAt_iff.mpr ⟨I.continuousAt, ?_⟩ simp only [mfld_simps] refine contDiffWithinAt_id.congr_of_eventuallyEq ?_ ?_ · exact Filter.eventuallyEq_of_mem self_mem_nhdsWithin fun x₂ => I.right_inv @@ -50,16 +50,16 @@ theorem contMDiffOn_model_symm : ContMDiffOn 𝓘(𝕜, E) I n I.symm (range I) /-- An atlas member is `C^n` for any `n`. -/ theorem contMDiffOn_of_mem_maximalAtlas (h : e ∈ maximalAtlas I M) : ContMDiffOn I I n e e.source := ContMDiffOn.of_le - ((contDiffWithinAt_localInvariantProp I I ∞).liftPropOn_of_mem_maximalAtlas - (contDiffWithinAtProp_id I) h) + ((contDiffWithinAt_localInvariantProp ∞).liftPropOn_of_mem_maximalAtlas + contDiffWithinAtProp_id h) le_top /-- The inverse of an atlas member is `C^n` for any `n`. -/ theorem contMDiffOn_symm_of_mem_maximalAtlas (h : e ∈ maximalAtlas I M) : ContMDiffOn I I n e.symm e.target := ContMDiffOn.of_le - ((contDiffWithinAt_localInvariantProp I I ∞).liftPropOn_symm_of_mem_maximalAtlas - (contDiffWithinAtProp_id I) h) + ((contDiffWithinAt_localInvariantProp ∞).liftPropOn_symm_of_mem_maximalAtlas + contDiffWithinAtProp_id h) le_top theorem contMDiffAt_of_mem_maximalAtlas (h : e ∈ maximalAtlas I M) (hx : x ∈ e.source) : @@ -71,10 +71,10 @@ theorem contMDiffAt_symm_of_mem_maximalAtlas {x : H} (h : e ∈ maximalAtlas I M (contMDiffOn_symm_of_mem_maximalAtlas h).contMDiffAt <| e.open_target.mem_nhds hx theorem contMDiffOn_chart : ContMDiffOn I I n (chartAt H x) (chartAt H x).source := - contMDiffOn_of_mem_maximalAtlas <| chart_mem_maximalAtlas I x + contMDiffOn_of_mem_maximalAtlas <| chart_mem_maximalAtlas x theorem contMDiffOn_chart_symm : ContMDiffOn I I n (chartAt H x).symm (chartAt H x).target := - contMDiffOn_symm_of_mem_maximalAtlas <| chart_mem_maximalAtlas I x + contMDiffOn_symm_of_mem_maximalAtlas <| chart_mem_maximalAtlas x theorem contMDiffAt_extend {x : M} (he : e ∈ maximalAtlas I M) (hx : x ∈ e.source) : ContMDiffAt I 𝓘(𝕜, E) n (e.extend I) x := @@ -82,7 +82,7 @@ theorem contMDiffAt_extend {x : M} (he : e ∈ maximalAtlas I M) (hx : x ∈ e.s theorem contMDiffAt_extChartAt' {x' : M} (h : x' ∈ (chartAt H x).source) : ContMDiffAt I 𝓘(𝕜, E) n (extChartAt I x) x' := - contMDiffAt_extend (chart_mem_maximalAtlas I x) h + contMDiffAt_extend (chart_mem_maximalAtlas x) h theorem contMDiffAt_extChartAt : ContMDiffAt I 𝓘(𝕜, E) n (extChartAt I x) x := contMDiffAt_extChartAt' <| mem_chart_source H x @@ -99,14 +99,13 @@ theorem contMDiffOn_extend_symm (he : e ∈ maximalAtlas I M) : theorem contMDiffOn_extChartAt_symm (x : M) : ContMDiffOn 𝓘(𝕜, E) I n (extChartAt I x).symm (extChartAt I x).target := by - convert contMDiffOn_extend_symm (chart_mem_maximalAtlas I x) + convert contMDiffOn_extend_symm (chart_mem_maximalAtlas (I := I) x) rw [extChartAt_target, I.image_eq] /-- An element of `contDiffGroupoid ⊤ I` is `C^n` for any `n`. -/ theorem contMDiffOn_of_mem_contDiffGroupoid {e' : PartialHomeomorph H H} (h : e' ∈ contDiffGroupoid ⊤ I) : ContMDiffOn I I n e' e'.source := - (contDiffWithinAt_localInvariantProp I I n).liftPropOn_of_mem_groupoid - (contDiffWithinAtProp_id I) h + (contDiffWithinAt_localInvariantProp n).liftPropOn_of_mem_groupoid contDiffWithinAtProp_id h end Atlas diff --git a/Mathlib/Geometry/Manifold/ContMDiff/Basic.lean b/Mathlib/Geometry/Manifold/ContMDiff/Basic.lean index b79ff42cc1ad9..a1064a541b84a 100644 --- a/Mathlib/Geometry/Manifold/ContMDiff/Basic.lean +++ b/Mathlib/Geometry/Manifold/ContMDiff/Basic.lean @@ -28,11 +28,11 @@ variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] -- declare the prerequisites for a charted space `M` over the pair `(E, H)`. {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] - (I : ModelWithCorners 𝕜 E H) {M : Type*} [TopologicalSpace M] + {I : ModelWithCorners 𝕜 E H} {M : Type*} [TopologicalSpace M] -- declare the prerequisites for a charted space `M'` over the pair `(E', H')`. {E' : Type*} [NormedAddCommGroup E'] [NormedSpace 𝕜 E'] {H' : Type*} [TopologicalSpace H'] - (I' : ModelWithCorners 𝕜 E' H') {M' : Type*} [TopologicalSpace M'] + {I' : ModelWithCorners 𝕜 E' H'} {M' : Type*} [TopologicalSpace M'] -- declare the prerequisites for a charted space `M''` over the pair `(E'', H'')`. {E'' : Type*} [NormedAddCommGroup E''] [NormedSpace 𝕜 E''] {H'' : Type*} [TopologicalSpace H''] @@ -43,7 +43,6 @@ variable [ChartedSpace H M] [ChartedSpace H' M'] [ChartedSpace H'' M''] -- declare functions, sets, points and smoothness indices {e : PartialHomeomorph M H} {e' : PartialHomeomorph M' H'} {f f₁ : M → M'} {s s₁ t : Set M} {x : M} {m n : ℕ∞} -variable {I I'} /-! ### Smoothness of the composition of smooth functions between manifolds -/ @@ -61,8 +60,8 @@ theorem ContMDiffWithinAt.comp {t : Set M'} {g : M' → M''} (x : M) rw [this] at hg have A : ∀ᶠ y in 𝓝[e.symm ⁻¹' s ∩ range I] e x, f (e.symm y) ∈ t ∧ f (e.symm y) ∈ e'.source := by simp only [e, ← map_extChartAt_nhdsWithin, eventually_map] - filter_upwards [hf.1.tendsto (extChartAt_source_mem_nhds I' (f x)), - inter_mem_nhdsWithin s (extChartAt_source_mem_nhds I x)] + filter_upwards [hf.1.tendsto (extChartAt_source_mem_nhds (I := I') (f x)), + inter_mem_nhdsWithin s (extChartAt_source_mem_nhds (I := I) x)] rintro x' (hfx' : f x' ∈ e'.source) ⟨hx's, hx'⟩ simp only [e.map_source hx', true_and, e.left_inv hx', st hx's, *] refine ((hg.2.comp _ (hf.2.mono inter_subset_right) inter_subset_left).mono_of_mem @@ -184,7 +183,7 @@ section id theorem contMDiff_id : ContMDiff I I n (id : M → M) := ContMDiff.of_le - ((contDiffWithinAt_localInvariantProp I I ∞).liftProp_id (contDiffWithinAtProp_id I)) le_top + ((contDiffWithinAt_localInvariantProp ∞).liftProp_id contDiffWithinAtProp_id) le_top theorem smooth_id : Smooth I I (id : M → M) := contMDiff_id @@ -310,7 +309,7 @@ open TopologicalSpace theorem contMdiffAt_subtype_iff {n : ℕ∞} {U : Opens M} {f : M → M'} {x : U} : ContMDiffAt I I' n (fun x : U ↦ f x) x ↔ ContMDiffAt I I' n f x := - ((contDiffWithinAt_localInvariantProp I I' n).liftPropAt_iff_comp_subtype_val _ _).symm + ((contDiffWithinAt_localInvariantProp n).liftPropAt_iff_comp_subtype_val _ _).symm theorem contMDiff_subtype_val {n : ℕ∞} {U : Opens M} : ContMDiff I I n (Subtype.val : U → M) := fun _ ↦ contMdiffAt_subtype_iff.mpr contMDiffAt_id @@ -330,7 +329,7 @@ theorem ContMDiff.extend_one [T2Space M] [One M'] {n : ℕ∞} {U : Opens M} {f theorem contMDiff_inclusion {n : ℕ∞} {U V : Opens M} (h : U ≤ V) : ContMDiff I I n (Opens.inclusion h : U → V) := by rintro ⟨x, hx : x ∈ U⟩ - apply (contDiffWithinAt_localInvariantProp I I n).liftProp_inclusion + apply (contDiffWithinAt_localInvariantProp n).liftProp_inclusion intro y dsimp only [ContDiffWithinAtProp, id_comp, preimage_univ] rw [Set.univ_inter] @@ -364,7 +363,7 @@ variable {e : M → H} (h : IsOpenEmbedding e) {n : WithTop ℕ} then `e` is smooth. -/ lemma contMDiff_isOpenEmbedding [Nonempty M] : haveI := h.singletonChartedSpace; ContMDiff I I n e := by - haveI := h.singleton_smoothManifoldWithCorners I + haveI := h.singleton_smoothManifoldWithCorners (I := I) rw [@contMDiff_iff _ _ _ _ _ _ _ _ _ _ h.singletonChartedSpace] use h.continuous intros x y @@ -376,7 +375,7 @@ lemma contMDiff_isOpenEmbedding [Nonempty M] : rw [h.toPartialHomeomorph_right_inv] · rw [I.right_inv] apply mem_of_subset_of_mem _ hz.1 - exact haveI := h.singletonChartedSpace; extChartAt_target_subset_range I x + exact haveI := h.singletonChartedSpace; extChartAt_target_subset_range (I := I) x · -- `hz` implies that `z ∈ range (I ∘ e)` have := hz.1 rw [@extChartAt_target _ _ _ _ _ _ _ _ _ _ h.singletonChartedSpace] at this @@ -388,13 +387,12 @@ lemma contMDiff_isOpenEmbedding [Nonempty M] : @[deprecated (since := "2024-10-18")] alias contMDiff_openEmbedding := contMDiff_isOpenEmbedding -variable {I I'} /-- If the `ChartedSpace` structure on a manifold `M` is given by an open embedding `e : M → H`, then the inverse of `e` is smooth. -/ lemma contMDiffOn_isOpenEmbedding_symm [Nonempty M] : haveI := h.singletonChartedSpace; ContMDiffOn I I n (IsOpenEmbedding.toPartialHomeomorph e h).symm (range e) := by - haveI := h.singleton_smoothManifoldWithCorners I + haveI := h.singleton_smoothManifoldWithCorners (I := I) rw [@contMDiffOn_iff] constructor · rw [← h.toPartialHomeomorph_target] @@ -410,7 +408,7 @@ lemma contMDiffOn_isOpenEmbedding_symm [Nonempty M] : exact hz.2.1 rw [h.toPartialHomeomorph_right_inv e this] apply I.right_inv - exact mem_of_subset_of_mem (extChartAt_target_subset_range _ _) hz.1 + exact mem_of_subset_of_mem (extChartAt_target_subset_range _) hz.1 @[deprecated (since := "2024-10-18")] alias contMDiffOn_openEmbedding_symm := contMDiffOn_isOpenEmbedding_symm diff --git a/Mathlib/Geometry/Manifold/ContMDiff/Defs.lean b/Mathlib/Geometry/Manifold/ContMDiff/Defs.lean index e52bf1aa65520..a9aa326b800c8 100644 --- a/Mathlib/Geometry/Manifold/ContMDiff/Defs.lean +++ b/Mathlib/Geometry/Manifold/ContMDiff/Defs.lean @@ -55,11 +55,11 @@ variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] -- Prerequisite typeclasses to say that `M` is a smooth manifold over the pair `(E, H)` {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] - (I : ModelWithCorners 𝕜 E H) {M : Type*} [TopologicalSpace M] [ChartedSpace H M] + {I : ModelWithCorners 𝕜 E H} {M : Type*} [TopologicalSpace M] [ChartedSpace H M] -- Prerequisite typeclasses to say that `M'` is a smooth manifold over the pair `(E', H')` {E' : Type*} [NormedAddCommGroup E'] [NormedSpace 𝕜 E'] {H' : Type*} [TopologicalSpace H'] - (I' : ModelWithCorners 𝕜 E' H') {M' : Type*} [TopologicalSpace M'] [ChartedSpace H' M'] + {I' : ModelWithCorners 𝕜 E' H'} {M' : Type*} [TopologicalSpace M'] [ChartedSpace H' M'] -- Prerequisite typeclasses to say that `M''` is a smooth manifold over the pair `(E'', H'')` {E'' : Type*} [NormedAddCommGroup E''] [NormedSpace 𝕜 E''] {H'' : Type*} [TopologicalSpace H''] @@ -68,6 +68,7 @@ variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {e : PartialHomeomorph M H} {e' : PartialHomeomorph M' H'} {f f₁ : M → M'} {s s₁ t : Set M} {x : M} {m n : ℕ∞} +variable (I I') in /-- Property in the model space of a model with corners of being `C^n` within at set at a point, when read in the model vector space. This property will be lifted to manifolds to define smooth functions between manifolds. -/ @@ -81,7 +82,7 @@ theorem contDiffWithinAtProp_self_source {f : E → H'} {s : Set E} {x : E} : theorem contDiffWithinAtProp_self {f : E → E'} {s : Set E} {x : E} : ContDiffWithinAtProp 𝓘(𝕜, E) 𝓘(𝕜, E') n f s x ↔ ContDiffWithinAt 𝕜 n f s x := - contDiffWithinAtProp_self_source 𝓘(𝕜, E') + contDiffWithinAtProp_self_source theorem contDiffWithinAtProp_self_target {f : H → E'} {s : Set H} {x : H} : ContDiffWithinAtProp I 𝓘(𝕜, E') n f s x ↔ @@ -143,17 +144,20 @@ theorem contDiffWithinAtProp_id (x : H) : ContDiffWithinAtProp I I n id univ x : · simp only [ModelWithCorners.right_inv I hy, mfld_simps] · simp only [mfld_simps] +variable (I I') in /-- A function is `n` times continuously differentiable within a set at a point in a manifold if it is continuous and it is `n` times continuously differentiable in this set around this point, when read in the preferred chart at this point. -/ def ContMDiffWithinAt (n : ℕ∞) (f : M → M') (s : Set M) (x : M) := LiftPropWithinAt (ContDiffWithinAtProp I I' n) f s x +variable (I I') in /-- Abbreviation for `ContMDiffWithinAt I I' ⊤ f s x`. See also documentation for `Smooth`. -/ abbrev SmoothWithinAt (f : M → M') (s : Set M) (x : M) := ContMDiffWithinAt I I' ⊤ f s x +variable (I I') in /-- A function is `n` times continuously differentiable at a point in a manifold if it is continuous and it is `n` times continuously differentiable around this point, when read in the preferred chart at this point. -/ @@ -167,26 +171,31 @@ theorem contMDiffAt_iff {n : ℕ∞} {f : M → M'} {x : M} : (extChartAt I x x) := liftPropAt_iff.trans <| by rw [ContDiffWithinAtProp, preimage_univ, univ_inter]; rfl +variable (I I') in /-- Abbreviation for `ContMDiffAt I I' ⊤ f x`. See also documentation for `Smooth`. -/ abbrev SmoothAt (f : M → M') (x : M) := ContMDiffAt I I' ⊤ f x +variable (I I') in /-- A function is `n` times continuously differentiable in a set of a manifold if it is continuous and, for any pair of points, it is `n` times continuously differentiable on this set in the charts around these points. -/ def ContMDiffOn (n : ℕ∞) (f : M → M') (s : Set M) := ∀ x ∈ s, ContMDiffWithinAt I I' n f s x +variable (I I') in /-- Abbreviation for `ContMDiffOn I I' ⊤ f s`. See also documentation for `Smooth`. -/ abbrev SmoothOn (f : M → M') (s : Set M) := ContMDiffOn I I' ⊤ f s +variable (I I') in /-- A function is `n` times continuously differentiable in a manifold if it is continuous and, for any pair of points, it is `n` times continuously differentiable in the charts around these points. -/ def ContMDiff (n : ℕ∞) (f : M → M') := ∀ x, ContMDiffAt I I' n f x +variable (I I') in /-- Abbreviation for `ContMDiff I I' ⊤ f`. Short note to work with these abbreviations: a lemma of the form `ContMDiffFoo.bar` will apply fine to an assumption `SmoothFoo` using dot notation or normal notation. @@ -197,8 +206,6 @@ This also applies to `SmoothAt`, `SmoothOn` and `SmoothWithinAt`. -/ abbrev Smooth (f : M → M') := ContMDiff I I' ⊤ f -variable {I I'} - /-! ### Deducing smoothness from higher smoothness -/ theorem ContMDiffWithinAt.of_le (hf : ContMDiffWithinAt I I' n f s x) (le : m ≤ n) : @@ -287,7 +294,7 @@ theorem contMDiffWithinAt_iff' : (extChartAt I x x) := by simp only [ContMDiffWithinAt, liftPropWithinAt_iff'] exact and_congr_right fun hc => contDiffWithinAt_congr_nhds <| - hc.nhdsWithin_extChartAt_symm_preimage_inter_range I I' + hc.nhdsWithin_extChartAt_symm_preimage_inter_range /-- One can reformulate smoothness within a set at a point as continuity within this set at this point, and smoothness in the corresponding extended chart in the target. -/ @@ -298,7 +305,7 @@ theorem contMDiffWithinAt_iff_target : have cont : ContinuousWithinAt f s x ∧ ContinuousWithinAt (extChartAt I' (f x) ∘ f) s x ↔ ContinuousWithinAt f s x := - and_iff_left_of_imp <| (continuousAt_extChartAt _ _).comp_continuousWithinAt + and_iff_left_of_imp <| (continuousAt_extChartAt _).comp_continuousWithinAt simp_rw [cont, ContDiffWithinAtProp, extChartAt, PartialHomeomorph.extend, PartialEquiv.coe_trans, ModelWithCorners.toPartialEquiv_coe, PartialHomeomorph.coe_coe, modelWithCornersSelf_coe, chartAt_self_eq, PartialHomeomorph.refl_apply, id_comp] @@ -332,9 +339,9 @@ theorem contMDiffWithinAt_iff_source_of_mem_maximalAtlas ContMDiffWithinAt I I' n f s x ↔ ContMDiffWithinAt 𝓘(𝕜, E) I' n (f ∘ (e.extend I).symm) ((e.extend I).symm ⁻¹' s ∩ range I) (e.extend I x) := by - have h2x := hx; rw [← e.extend_source I] at h2x + have h2x := hx; rw [← e.extend_source (I := I)] at h2x simp_rw [ContMDiffWithinAt, - (contDiffWithinAt_localInvariantProp I I' n).liftPropWithinAt_indep_chart_source he hx, + (contDiffWithinAt_localInvariantProp n).liftPropWithinAt_indep_chart_source he hx, StructureGroupoid.liftPropWithinAt_self_source, e.extend_symm_continuousWithinAt_comp_right_iff, contDiffWithinAtProp_self_source, ContDiffWithinAtProp, Function.comp, e.left_inv hx, (e.extend I).left_inv h2x] @@ -345,7 +352,7 @@ theorem contMDiffWithinAt_iff_source_of_mem_source ContMDiffWithinAt I I' n f s x' ↔ ContMDiffWithinAt 𝓘(𝕜, E) I' n (f ∘ (extChartAt I x).symm) ((extChartAt I x).symm ⁻¹' s ∩ range I) (extChartAt I x x') := - contMDiffWithinAt_iff_source_of_mem_maximalAtlas (chart_mem_maximalAtlas I x) hx' + contMDiffWithinAt_iff_source_of_mem_maximalAtlas (chart_mem_maximalAtlas x) hx' theorem contMDiffAt_iff_source_of_mem_source [SmoothManifoldWithCorners I M] {x' : M} (hx' : x' ∈ (chartAt H x).source) : @@ -358,14 +365,14 @@ theorem contMDiffWithinAt_iff_target_of_mem_source ContMDiffWithinAt I I' n f s x ↔ ContinuousWithinAt f s x ∧ ContMDiffWithinAt I 𝓘(𝕜, E') n (extChartAt I' y ∘ f) s x := by simp_rw [ContMDiffWithinAt] - rw [(contDiffWithinAt_localInvariantProp I I' n).liftPropWithinAt_indep_chart_target - (chart_mem_maximalAtlas I' y) hy, + rw [(contDiffWithinAt_localInvariantProp n).liftPropWithinAt_indep_chart_target + (chart_mem_maximalAtlas y) hy, and_congr_right] intro hf simp_rw [StructureGroupoid.liftPropWithinAt_self_target] simp_rw [((chartAt H' y).continuousAt hy).comp_continuousWithinAt hf] - rw [← extChartAt_source I'] at hy - simp_rw [(continuousAt_extChartAt' I' hy).comp_continuousWithinAt hf] + rw [← extChartAt_source (I := I')] at hy + simp_rw [(continuousAt_extChartAt' hy).comp_continuousWithinAt hf] rfl theorem contMDiffAt_iff_target_of_mem_source @@ -383,7 +390,7 @@ theorem contMDiffWithinAt_iff_of_mem_maximalAtlas {x : M} (he : e ∈ maximalAtl ContinuousWithinAt f s x ∧ ContDiffWithinAt 𝕜 n (e'.extend I' ∘ f ∘ (e.extend I).symm) ((e.extend I).symm ⁻¹' s ∩ range I) (e.extend I x) := - (contDiffWithinAt_localInvariantProp I I' n).liftPropWithinAt_indep_chart he hx he' hy + (contDiffWithinAt_localInvariantProp n).liftPropWithinAt_indep_chart he hx he' hy /-- An alternative formulation of `contMDiffWithinAt_iff_of_mem_maximalAtlas` if the set if `s` lies in `e.source`. -/ @@ -395,7 +402,7 @@ theorem contMDiffWithinAt_iff_image {x : M} (he : e ∈ maximalAtlas I M) (e.extend I x) := by rw [contMDiffWithinAt_iff_of_mem_maximalAtlas he he' hx hy, and_congr_right_iff] refine fun _ => contDiffWithinAt_congr_nhds ?_ - simp_rw [nhdsWithin_eq_iff_eventuallyEq, e.extend_symm_preimage_inter_range_eventuallyEq I hs hx] + simp_rw [nhdsWithin_eq_iff_eventuallyEq, e.extend_symm_preimage_inter_range_eventuallyEq hs hx] /-- One can reformulate smoothness within a set at a point as continuity within this set at this point, and smoothness in any chart containing that point. -/ @@ -405,8 +412,8 @@ theorem contMDiffWithinAt_iff_of_mem_source {x' : M} {y : M'} (hx : x' ∈ (char ContinuousWithinAt f s x' ∧ ContDiffWithinAt 𝕜 n (extChartAt I' y ∘ f ∘ (extChartAt I x).symm) ((extChartAt I x).symm ⁻¹' s ∩ range I) (extChartAt I x x') := - contMDiffWithinAt_iff_of_mem_maximalAtlas (chart_mem_maximalAtlas _ x) - (chart_mem_maximalAtlas _ y) hx hy + contMDiffWithinAt_iff_of_mem_maximalAtlas (chart_mem_maximalAtlas x) + (chart_mem_maximalAtlas y) hx hy theorem contMDiffWithinAt_iff_of_mem_source' {x' : M} {y : M'} (hx : x' ∈ (chartAt H x).source) (hy : f x' ∈ (chartAt H' y).source) : @@ -421,9 +428,9 @@ theorem contMDiffWithinAt_iff_of_mem_source' {x' : M} {y : M'} (hx : x' ∈ (cha rw [and_congr_right_iff] set e := extChartAt I x; set e' := extChartAt I' (f x) refine fun hc => contDiffWithinAt_congr_nhds ?_ - rw [← e.image_source_inter_eq', ← map_extChartAt_nhdsWithin_eq_image' I hx, ← - map_extChartAt_nhdsWithin' I hx, inter_comm, nhdsWithin_inter_of_mem] - exact hc (extChartAt_source_mem_nhds' _ hy) + rw [← e.image_source_inter_eq', ← map_extChartAt_nhdsWithin_eq_image' hx, + ← map_extChartAt_nhdsWithin' hx, inter_comm, nhdsWithin_inter_of_mem] + exact hc (extChartAt_source_mem_nhds' hy) theorem contMDiffAt_iff_of_mem_source {x' : M} {y : M'} (hx : x' ∈ (chartAt H x).source) (hy : f x' ∈ (chartAt H' y).source) : @@ -447,7 +454,7 @@ theorem contMDiffOn_iff_of_mem_maximalAtlas' (he : e ∈ maximalAtlas I M) ContMDiffOn I I' n f s ↔ ContDiffOn 𝕜 n (e'.extend I' ∘ f ∘ (e.extend I).symm) (e.extend I '' s) := (contMDiffOn_iff_of_mem_maximalAtlas he he' hs h2s).trans <| and_iff_right_of_imp fun h ↦ - (e.continuousOn_writtenInExtend_iff _ _ hs h2s).1 h.continuousOn + (e.continuousOn_writtenInExtend_iff hs h2s).1 h.continuousOn /-- If the set where you want `f` to be smooth lies entirely in a single chart, and `f` maps it into a single chart, the smoothness of `f` on that set can be expressed by purely looking in @@ -459,7 +466,7 @@ theorem contMDiffOn_iff_of_subset_source {x : M} {y : M'} (hs : s ⊆ (chartAt H ContMDiffOn I I' n f s ↔ ContinuousOn f s ∧ ContDiffOn 𝕜 n (extChartAt I' y ∘ f ∘ (extChartAt I x).symm) (extChartAt I x '' s) := - contMDiffOn_iff_of_mem_maximalAtlas (chart_mem_maximalAtlas I x) (chart_mem_maximalAtlas I' y) hs + contMDiffOn_iff_of_mem_maximalAtlas (chart_mem_maximalAtlas x) (chart_mem_maximalAtlas y) hs h2s /-- If the set where you want `f` to be smooth lies entirely in a single chart, and `f` maps it @@ -472,8 +479,8 @@ theorem contMDiffOn_iff_of_subset_source' {x : M} {y : M'} (hs : s ⊆ (extChart ContMDiffOn I I' n f s ↔ ContDiffOn 𝕜 n (extChartAt I' y ∘ f ∘ (extChartAt I x).symm) (extChartAt I x '' s) := by rw [extChartAt_source] at hs h2s - exact contMDiffOn_iff_of_mem_maximalAtlas' (chart_mem_maximalAtlas I x) - (chart_mem_maximalAtlas I' y) hs h2s + exact contMDiffOn_iff_of_mem_maximalAtlas' (chart_mem_maximalAtlas x) + (chart_mem_maximalAtlas y) hs h2s /-- One can reformulate smoothness on a set as continuity on this set, and smoothness in any extended chart. -/ @@ -497,7 +504,7 @@ theorem contMDiffOn_iff : · simp only [w, hz, mfld_simps] · mfld_set_tac · rintro ⟨hcont, hdiff⟩ x hx - refine (contDiffWithinAt_localInvariantProp I I' n).liftPropWithinAt_iff.mpr ?_ + refine (contDiffWithinAt_localInvariantProp n).liftPropWithinAt_iff.mpr ?_ refine ⟨hcont x hx, ?_⟩ dsimp [ContDiffWithinAtProp] convert hdiff x (f x) (extChartAt I x x) (by simp only [hx, mfld_simps]) using 1 @@ -632,7 +639,7 @@ theorem contMDiffWithinAt_iff_nat : theorem ContMDiffWithinAt.mono_of_mem (hf : ContMDiffWithinAt I I' n f s x) (hts : s ∈ 𝓝[t] x) : ContMDiffWithinAt I I' n f t x := StructureGroupoid.LocalInvariantProp.liftPropWithinAt_mono_of_mem - (contDiffWithinAtProp_mono_of_mem I I' n) hf hts + (contDiffWithinAtProp_mono_of_mem n) hf hts theorem ContMDiffWithinAt.mono (hf : ContMDiffWithinAt I I' n f s x) (hts : t ⊆ s) : ContMDiffWithinAt I I' n f t x := @@ -647,7 +654,7 @@ theorem contMDiffWithinAt_insert_self : ContMDiffWithinAt I I' n f (insert x s) x ↔ ContMDiffWithinAt I I' n f s x := by simp only [contMDiffWithinAt_iff, continuousWithinAt_insert_self] refine Iff.rfl.and <| (contDiffWithinAt_congr_nhds ?_).trans contDiffWithinAt_insert_self - simp only [← map_extChartAt_nhdsWithin I, nhdsWithin_insert, Filter.map_sup, Filter.map_pure] + simp only [← map_extChartAt_nhdsWithin, nhdsWithin_insert, Filter.map_sup, Filter.map_pure] alias ⟨ContMDiffWithinAt.of_insert, _⟩ := contMDiffWithinAt_insert_self @@ -674,15 +681,15 @@ theorem Smooth.smoothOn (hf : Smooth I I' f) : SmoothOn I I' f s := theorem contMDiffWithinAt_inter' (ht : t ∈ 𝓝[s] x) : ContMDiffWithinAt I I' n f (s ∩ t) x ↔ ContMDiffWithinAt I I' n f s x := - (contDiffWithinAt_localInvariantProp I I' n).liftPropWithinAt_inter' ht + (contDiffWithinAt_localInvariantProp n).liftPropWithinAt_inter' ht theorem contMDiffWithinAt_inter (ht : t ∈ 𝓝 x) : ContMDiffWithinAt I I' n f (s ∩ t) x ↔ ContMDiffWithinAt I I' n f s x := - (contDiffWithinAt_localInvariantProp I I' n).liftPropWithinAt_inter ht + (contDiffWithinAt_localInvariantProp n).liftPropWithinAt_inter ht theorem ContMDiffWithinAt.contMDiffAt (h : ContMDiffWithinAt I I' n f s x) (ht : s ∈ 𝓝 x) : ContMDiffAt I I' n f x := - (contDiffWithinAt_localInvariantProp I I' n).liftPropAt_of_liftPropWithinAt h ht + (contDiffWithinAt_localInvariantProp n).liftPropAt_of_liftPropWithinAt h ht theorem SmoothWithinAt.smoothAt (h : SmoothWithinAt I I' f s x) (ht : s ∈ 𝓝 x) : SmoothAt I I' f x := @@ -704,7 +711,7 @@ theorem contMDiffOn_iff_source_of_mem_maximalAtlas [SmoothManifoldWithCorners I rw [contMDiffWithinAt_iff_source_of_mem_maximalAtlas he (hs hx)] apply contMDiffWithinAt_congr_nhds simp_rw [nhdsWithin_eq_iff_eventuallyEq, - e.extend_symm_preimage_inter_range_eventuallyEq I hs (hs hx)] + e.extend_symm_preimage_inter_range_eventuallyEq hs (hs hx)] -- Porting note: didn't compile; fixed by golfing the proof and moving parts to lemmas /-- A function is `C^n` within a set at a point, for `n : ℕ`, if and only if it is `C^n` on @@ -723,13 +730,13 @@ theorem contMDiffWithinAt_iff_contMDiffOn_nhds rcases (contMDiffWithinAt_iff'.1 h).2.contDiffOn le_rfl with ⟨v, hmem, hsub, hv⟩ have hxs' : extChartAt I x x ∈ (extChartAt I x).target ∩ (extChartAt I x).symm ⁻¹' (s ∩ f ⁻¹' (extChartAt I' (f x)).source) := - ⟨(extChartAt I x).map_source (mem_extChartAt_source _ _), by rwa [extChartAt_to_inv], by + ⟨(extChartAt I x).map_source (mem_extChartAt_source _), by rwa [extChartAt_to_inv], by rw [extChartAt_to_inv]; apply mem_extChartAt_source⟩ rw [insert_eq_of_mem hxs'] at hmem hsub -- Then `(extChartAt I x).symm '' v` is the neighborhood we are looking for. refine ⟨(extChartAt I x).symm '' v, ?_, ?_⟩ - · rw [← map_extChartAt_symm_nhdsWithin I, - h.1.nhdsWithin_extChartAt_symm_preimage_inter_range I I'] + · rw [← map_extChartAt_symm_nhdsWithin (I := I), + h.1.nhdsWithin_extChartAt_symm_preimage_inter_range (I := I) (I' := I')] exact image_mem_map hmem · have hv₁ : (extChartAt I x).symm '' v ⊆ (extChartAt I x).source := image_subset_iff.2 fun y hy ↦ (extChartAt I x).map_target (hsub hy).1 @@ -761,35 +768,35 @@ theorem contMDiffAt_iff_contMDiffAt_nhds theorem ContMDiffWithinAt.congr (h : ContMDiffWithinAt I I' n f s x) (h₁ : ∀ y ∈ s, f₁ y = f y) (hx : f₁ x = f x) : ContMDiffWithinAt I I' n f₁ s x := - (contDiffWithinAt_localInvariantProp I I' n).liftPropWithinAt_congr h h₁ hx + (contDiffWithinAt_localInvariantProp n).liftPropWithinAt_congr h h₁ hx theorem contMDiffWithinAt_congr (h₁ : ∀ y ∈ s, f₁ y = f y) (hx : f₁ x = f x) : ContMDiffWithinAt I I' n f₁ s x ↔ ContMDiffWithinAt I I' n f s x := - (contDiffWithinAt_localInvariantProp I I' n).liftPropWithinAt_congr_iff h₁ hx + (contDiffWithinAt_localInvariantProp n).liftPropWithinAt_congr_iff h₁ hx theorem ContMDiffWithinAt.congr_of_eventuallyEq (h : ContMDiffWithinAt I I' n f s x) (h₁ : f₁ =ᶠ[𝓝[s] x] f) (hx : f₁ x = f x) : ContMDiffWithinAt I I' n f₁ s x := - (contDiffWithinAt_localInvariantProp I I' n).liftPropWithinAt_congr_of_eventuallyEq h h₁ hx + (contDiffWithinAt_localInvariantProp n).liftPropWithinAt_congr_of_eventuallyEq h h₁ hx theorem Filter.EventuallyEq.contMDiffWithinAt_iff (h₁ : f₁ =ᶠ[𝓝[s] x] f) (hx : f₁ x = f x) : ContMDiffWithinAt I I' n f₁ s x ↔ ContMDiffWithinAt I I' n f s x := - (contDiffWithinAt_localInvariantProp I I' n).liftPropWithinAt_congr_iff_of_eventuallyEq h₁ hx + (contDiffWithinAt_localInvariantProp n).liftPropWithinAt_congr_iff_of_eventuallyEq h₁ hx theorem ContMDiffAt.congr_of_eventuallyEq (h : ContMDiffAt I I' n f x) (h₁ : f₁ =ᶠ[𝓝 x] f) : ContMDiffAt I I' n f₁ x := - (contDiffWithinAt_localInvariantProp I I' n).liftPropAt_congr_of_eventuallyEq h h₁ + (contDiffWithinAt_localInvariantProp n).liftPropAt_congr_of_eventuallyEq h h₁ theorem Filter.EventuallyEq.contMDiffAt_iff (h₁ : f₁ =ᶠ[𝓝 x] f) : ContMDiffAt I I' n f₁ x ↔ ContMDiffAt I I' n f x := - (contDiffWithinAt_localInvariantProp I I' n).liftPropAt_congr_iff_of_eventuallyEq h₁ + (contDiffWithinAt_localInvariantProp n).liftPropAt_congr_iff_of_eventuallyEq h₁ theorem ContMDiffOn.congr (h : ContMDiffOn I I' n f s) (h₁ : ∀ y ∈ s, f₁ y = f y) : ContMDiffOn I I' n f₁ s := - (contDiffWithinAt_localInvariantProp I I' n).liftPropOn_congr h h₁ + (contDiffWithinAt_localInvariantProp n).liftPropOn_congr h h₁ theorem contMDiffOn_congr (h₁ : ∀ y ∈ s, f₁ y = f y) : ContMDiffOn I I' n f₁ s ↔ ContMDiffOn I I' n f s := - (contDiffWithinAt_localInvariantProp I I' n).liftPropOn_congr_iff h₁ + (contDiffWithinAt_localInvariantProp n).liftPropOn_congr_iff h₁ theorem ContMDiffOn.congr_mono (hf : ContMDiffOn I I' n f s) (h₁ : ∀ y ∈ s₁, f₁ y = f y) (hs : s₁ ⊆ s) : ContMDiffOn I I' n f₁ s₁ := @@ -801,8 +808,8 @@ theorem ContMDiffOn.congr_mono (hf : ContMDiffOn I I' n f s) (h₁ : ∀ y ∈ s /-- Being `C^n` is a local property. -/ theorem contMDiffOn_of_locally_contMDiffOn (h : ∀ x ∈ s, ∃ u, IsOpen u ∧ x ∈ u ∧ ContMDiffOn I I' n f (s ∩ u)) : ContMDiffOn I I' n f s := - (contDiffWithinAt_localInvariantProp I I' n).liftPropOn_of_locally_liftPropOn h + (contDiffWithinAt_localInvariantProp n).liftPropOn_of_locally_liftPropOn h theorem contMDiff_of_locally_contMDiffOn (h : ∀ x, ∃ u, IsOpen u ∧ x ∈ u ∧ ContMDiffOn I I' n f u) : ContMDiff I I' n f := - (contDiffWithinAt_localInvariantProp I I' n).liftProp_of_locally_liftPropOn h + (contDiffWithinAt_localInvariantProp n).liftProp_of_locally_liftPropOn h diff --git a/Mathlib/Geometry/Manifold/ContMDiff/Product.lean b/Mathlib/Geometry/Manifold/ContMDiff/Product.lean index 5c76e05789139..fcd0a708a6a66 100644 --- a/Mathlib/Geometry/Manifold/ContMDiff/Product.lean +++ b/Mathlib/Geometry/Manifold/ContMDiff/Product.lean @@ -23,11 +23,11 @@ variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] -- declare a charted space `M` over the pair `(E, H)`. {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] - (I : ModelWithCorners 𝕜 E H) {M : Type*} [TopologicalSpace M] [ChartedSpace H M] + {I : ModelWithCorners 𝕜 E H} {M : Type*} [TopologicalSpace M] [ChartedSpace H M] -- declare a charted space `M'` over the pair `(E', H')`. {E' : Type*} [NormedAddCommGroup E'] [NormedSpace 𝕜 E'] {H' : Type*} [TopologicalSpace H'] - (I' : ModelWithCorners 𝕜 E' H') {M' : Type*} [TopologicalSpace M'] [ChartedSpace H' M'] + {I' : ModelWithCorners 𝕜 E' H'} {M' : Type*} [TopologicalSpace M'] [ChartedSpace H' M'] -- declare a charted space `N` over the pair `(F, G)`. {F : Type*} [NormedAddCommGroup F] [NormedSpace 𝕜 F] {G : Type*} [TopologicalSpace G] @@ -38,7 +38,6 @@ variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {J' : ModelWithCorners 𝕜 F' G'} {N' : Type*} [TopologicalSpace N'] [ChartedSpace G' N'] -- declare functions, sets, points and smoothness indices {f : M → M'} {s : Set M} {x : M} {n : ℕ∞} -variable {I I'} section ProdMk @@ -129,7 +128,7 @@ theorem contMDiffWithinAt_fst {s : Set (M × N)} {p : M × N} : rw [contMDiffWithinAt_iff'] refine ⟨continuousWithinAt_fst, contDiffWithinAt_fst.congr (fun y hy => ?_) ?_⟩ · exact (extChartAt I p.1).right_inv ⟨hy.1.1.1, hy.1.2.1⟩ - · exact (extChartAt I p.1).right_inv <| (extChartAt I p.1).map_source (mem_extChartAt_source _ _) + · exact (extChartAt I p.1).right_inv <| (extChartAt I p.1).map_source (mem_extChartAt_source _) theorem ContMDiffWithinAt.fst {f : N → M × M'} {s : Set N} {x : N} (hf : ContMDiffWithinAt J (I.prod I') n f s x) : @@ -185,7 +184,7 @@ theorem contMDiffWithinAt_snd {s : Set (M × N)} {p : M × N} : rw [contMDiffWithinAt_iff'] refine ⟨continuousWithinAt_snd, contDiffWithinAt_snd.congr (fun y hy => ?_) ?_⟩ · exact (extChartAt J p.2).right_inv ⟨hy.1.1.2, hy.1.2.2⟩ - · exact (extChartAt J p.2).right_inv <| (extChartAt J p.2).map_source (mem_extChartAt_source _ _) + · exact (extChartAt J p.2).right_inv <| (extChartAt J p.2).map_source (mem_extChartAt_source _) theorem ContMDiffWithinAt.snd {f : N → M × M'} {s : Set N} {x : N} (hf : ContMDiffWithinAt J (I.prod I') n f s x) : diff --git a/Mathlib/Geometry/Manifold/ContMDiffMFDeriv.lean b/Mathlib/Geometry/Manifold/ContMDiffMFDeriv.lean index efcc63b456d66..0f305e578a8e3 100644 --- a/Mathlib/Geometry/Manifold/ContMDiffMFDeriv.lean +++ b/Mathlib/Geometry/Manifold/ContMDiffMFDeriv.lean @@ -76,12 +76,12 @@ protected theorem ContMDiffAt.mfderiv {x₀ : N} (f : N → M → M') (g : N → x₀ := by have h4f : ContinuousAt (fun x => f x (g x)) x₀ := ContinuousAt.comp_of_eq hf.continuousAt (continuousAt_id.prod hg.continuousAt) rfl - have h4f := h4f.preimage_mem_nhds (extChartAt_source_mem_nhds I' (f x₀ (g x₀))) + have h4f := h4f.preimage_mem_nhds (extChartAt_source_mem_nhds (I := I') (f x₀ (g x₀))) have h3f := contMDiffAt_iff_contMDiffAt_nhds.mp (hf.of_le <| (self_le_add_left 1 m).trans hmn) have h2f : ∀ᶠ x₂ in 𝓝 x₀, ContMDiffAt I I' 1 (f x₂) (g x₂) := by refine ((continuousAt_id.prod hg.continuousAt).tendsto.eventually h3f).mono fun x hx => ?_ exact hx.comp (g x) (contMDiffAt_const.prod_mk contMDiffAt_id) - have h2g := hg.continuousAt.preimage_mem_nhds (extChartAt_source_mem_nhds I (g x₀)) + have h2g := hg.continuousAt.preimage_mem_nhds (extChartAt_source_mem_nhds (I := I) (g x₀)) have : ContDiffWithinAt 𝕜 m (fun x => @@ -130,9 +130,9 @@ protected theorem ContMDiffAt.mfderiv {x₀ : N} (f : N → M → M') (g : N → rw [(extChartAt I (g x₂)).left_inv hx, (extChartAt I' (f x₂ (g x₂))).left_inv h2x] refine Filter.EventuallyEq.fderivWithin_eq_nhds ?_ refine eventually_of_mem (inter_mem ?_ ?_) this - · exact extChartAt_preimage_mem_nhds' _ hx₂ (extChartAt_source_mem_nhds I (g x₂)) - · refine extChartAt_preimage_mem_nhds' _ hx₂ ?_ - exact h2x₂.continuousAt.preimage_mem_nhds (extChartAt_source_mem_nhds _ _) + · exact extChartAt_preimage_mem_nhds' hx₂ (extChartAt_source_mem_nhds (g x₂)) + · refine extChartAt_preimage_mem_nhds' hx₂ ?_ + exact h2x₂.continuousAt.preimage_mem_nhds (extChartAt_source_mem_nhds _) /- The conclusion is equal to the following, when unfolding coord_change of `tangentBundleCore` -/ -- Porting note: added @@ -153,19 +153,19 @@ protected theorem ContMDiffAt.mfderiv {x₀ : N} (f : N → M → M') (g : N → intro x₂ hx₂ h2x₂ h3x₂ symm rw [(h2x₂.mdifferentiableAt le_rfl).mfderiv] - have hI := (contDiffWithinAt_ext_coord_change I (g x₂) (g x₀) <| + have hI := (contDiffWithinAt_ext_coord_change (I := I) (g x₂) (g x₀) <| PartialEquiv.mem_symm_trans_source _ hx₂ <| - mem_extChartAt_source I (g x₂)).differentiableWithinAt le_top + mem_extChartAt_source (g x₂)).differentiableWithinAt le_top have hI' := - (contDiffWithinAt_ext_coord_change I' (f x₀ (g x₀)) (f x₂ (g x₂)) <| - PartialEquiv.mem_symm_trans_source _ (mem_extChartAt_source I' (f x₂ (g x₂))) + (contDiffWithinAt_ext_coord_change (f x₀ (g x₀)) (f x₂ (g x₂)) <| + PartialEquiv.mem_symm_trans_source _ (mem_extChartAt_source (f x₂ (g x₂))) h3x₂).differentiableWithinAt le_top have h3f := (h2x₂.mdifferentiableAt le_rfl).differentiableWithinAt_writtenInExtChartAt refine fderivWithin.comp₃ _ hI' h3f hI ?_ ?_ ?_ ?_ (I.uniqueDiffOn _ <| mem_range_self _) · exact fun x _ => mem_range_self _ · exact fun x _ => mem_range_self _ · simp_rw [writtenInExtChartAt, Function.comp_apply, - (extChartAt I (g x₂)).left_inv (mem_extChartAt_source I (g x₂))] + (extChartAt I (g x₂)).left_inv (mem_extChartAt_source (g x₂))] · simp_rw [Function.comp_apply, (extChartAt I (g x₀)).left_inv hx₂] refine this.congr_of_eventuallyEq ?_ filter_upwards [h2g, h4f] with x hx h2x @@ -223,12 +223,12 @@ theorem ContMDiffOn.continuousOn_tangentMapWithin_aux {f : H → H'} {s : Set H} (f p.fst, (fderivWithin 𝕜 (writtenInExtChartAt I I' p.fst f) (I.symm ⁻¹' s ∩ range I) ((extChartAt I p.fst) p.fst) : E →L[𝕜] E') p.snd)) (Prod.fst ⁻¹' s) by - have A := (tangentBundleModelSpaceHomeomorph H I).continuous + have A := (tangentBundleModelSpaceHomeomorph I).continuous rw [continuous_iff_continuousOn_univ] at A have B := - ((tangentBundleModelSpaceHomeomorph H' I').symm.continuous.comp_continuousOn h).comp' A + ((tangentBundleModelSpaceHomeomorph I').symm.continuous.comp_continuousOn h).comp' A have : - univ ∩ tangentBundleModelSpaceHomeomorph H I ⁻¹' (Prod.fst ⁻¹' s) = + univ ∩ tangentBundleModelSpaceHomeomorph I ⁻¹' (Prod.fst ⁻¹' s) = π E (TangentSpace I) ⁻¹' s := by ext ⟨x, v⟩; simp only [mfld_simps] rw [this] at B @@ -390,7 +390,7 @@ theorem ContMDiffOn.contMDiffOn_tangentMapWithin apply UniqueMDiffOn.inter _ l.open_source rw [ho, inter_comm] exact hs.inter o_open - have U'l : UniqueMDiffOn I s'l := U'.uniqueMDiffOn_preimage (mdifferentiable_chart _ _) + have U'l : UniqueMDiffOn I s'l := U'.uniqueMDiffOn_preimage (mdifferentiable_chart _) have diff_f : ContMDiffOn I I' n f s' := hf.mono (by unfold_let s'; mfld_set_tac) have diff_r : ContMDiffOn I' I' n r r.source := contMDiffOn_chart have diff_rf : ContMDiffOn I I' n (r ∘ f) s' := by @@ -461,7 +461,7 @@ theorem ContMDiffOn.contMDiffOn_tangentMapWithin tangentMap I I l.symm (il.symm (Dl q)) := by refine tangentMapWithin_eq_tangentMap U'lq ?_ -- Porting note: the arguments below were underscores. - refine mdifferentiableAt_atlas_symm I (chart_mem_atlas H (TotalSpace.proj p)) ?_ + refine mdifferentiableAt_atlas_symm (chart_mem_atlas H (TotalSpace.proj p)) ?_ simp only [Dl, il, hq, mfld_simps] rw [this, tangentMap_chart_symm, hDl] · simp only [il, hq, mfld_simps] @@ -487,7 +487,7 @@ theorem ContMDiffOn.contMDiffOn_tangentMapWithin apply tangentMapWithin_eq_tangentMap · apply r.open_source.uniqueMDiffWithinAt _ simp [hq] - · exact mdifferentiableAt_atlas I' (chart_mem_atlas H' (f p.proj)) hq.1.1 + · exact mdifferentiableAt_atlas (chart_mem_atlas H' (f p.proj)) hq.1.1 have : f p.proj = (tangentMapWithin I I' f s p).1 := rfl rw [A] dsimp [Dr, ir, s', r, l] @@ -538,8 +538,6 @@ end tangentMap namespace TangentBundle -variable (I M) - open Bundle /-- The derivative of the zero section of the tangent bundle maps `⟨x, v⟩` to `⟨⟨x, 0⟩, ⟨v, 0⟩⟩`. diff --git a/Mathlib/Geometry/Manifold/ContMDiffMap.lean b/Mathlib/Geometry/Manifold/ContMDiffMap.lean index 9eac199f01e6d..8566926007fa4 100644 --- a/Mathlib/Geometry/Manifold/ContMDiffMap.lean +++ b/Mathlib/Geometry/Manifold/ContMDiffMap.lean @@ -14,8 +14,8 @@ bundled maps. variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {E' : Type*} [NormedAddCommGroup E'] [NormedSpace 𝕜 E'] {H : Type*} - [TopologicalSpace H] {H' : Type*} [TopologicalSpace H'] (I : ModelWithCorners 𝕜 E H) - (I' : ModelWithCorners 𝕜 E' H') (M : Type*) [TopologicalSpace M] [ChartedSpace H M] (M' : Type*) + [TopologicalSpace H] {H' : Type*} [TopologicalSpace H'] {I : ModelWithCorners 𝕜 E H} + {I' : ModelWithCorners 𝕜 E' H'} (M : Type*) [TopologicalSpace M] [ChartedSpace H M] (M' : Type*) [TopologicalSpace M'] [ChartedSpace H' M'] {E'' : Type*} [NormedAddCommGroup E''] [NormedSpace 𝕜 E''] {H'' : Type*} [TopologicalSpace H''] {I'' : ModelWithCorners 𝕜 E'' H''} {M'' : Type*} [TopologicalSpace M''] [ChartedSpace H'' M''] @@ -24,10 +24,12 @@ variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedAddCom [NormedAddCommGroup F] [NormedSpace 𝕜 F] {G : Type*} [TopologicalSpace G] {J : ModelWithCorners 𝕜 F G} {N : Type*} [TopologicalSpace N] [ChartedSpace G N] (n : ℕ∞) +variable (I I') in /-- Bundled `n` times continuously differentiable maps. -/ def ContMDiffMap := { f : M → M' // ContMDiff I I' n f } +variable (I I') in /-- Bundled smooth maps. -/ abbrev SmoothMap := ContMDiffMap I I' M M' ⊤ @@ -43,7 +45,7 @@ open scoped Manifold namespace ContMDiffMap -variable {I} {I'} {M} {M'} {n} +variable {M} {M'} {n} instance instFunLike : FunLike C^n⟮I, M; I', M'⟯ M M' where coe := Subtype.val diff --git a/Mathlib/Geometry/Manifold/Diffeomorph.lean b/Mathlib/Geometry/Manifold/Diffeomorph.lean index fa71b1fc3e365..88e23753a1297 100644 --- a/Mathlib/Geometry/Manifold/Diffeomorph.lean +++ b/Mathlib/Geometry/Manifold/Diffeomorph.lean @@ -497,14 +497,14 @@ def toTransDiffeomorph (e : E ≃ₘ[𝕜] F) : M ≃ₘ⟮I, I.transDiffeomorph · simp only [Equiv.coe_refl, id, (· ∘ ·), I.coe_extChartAt_transDiffeomorph, (extChartAt I x).right_inv hy.1] · exact - ⟨(extChartAt I x).map_source (mem_extChartAt_source I x), trivial, by simp only [mfld_simps]⟩ + ⟨(extChartAt I x).map_source (mem_extChartAt_source x), trivial, by simp only [mfld_simps]⟩ contMDiff_invFun x := by refine contMDiffWithinAt_iff'.2 ⟨continuousWithinAt_id, ?_⟩ refine e.symm.contDiff.contDiffWithinAt.congr' (fun y hy => ?_) ?_ · simp only [mem_inter_iff, I.extChartAt_transDiffeomorph_target] at hy simp only [Equiv.coe_refl, Equiv.refl_symm, id, (· ∘ ·), I.coe_extChartAt_transDiffeomorph_symm, (extChartAt I x).right_inv hy.1] - exact ⟨(extChartAt _ x).map_source (mem_extChartAt_source _ x), trivial, by + exact ⟨(extChartAt _ x).map_source (mem_extChartAt_source x), trivial, by simp only [e.symm_apply_apply, Equiv.refl_symm, Equiv.coe_refl, mfld_simps]⟩ variable {I M} diff --git a/Mathlib/Geometry/Manifold/Instances/Real.lean b/Mathlib/Geometry/Manifold/Instances/Real.lean index f3df629fd809b..9b54f975c0526 100644 --- a/Mathlib/Geometry/Manifold/Instances/Real.lean +++ b/Mathlib/Geometry/Manifold/Instances/Real.lean @@ -323,7 +323,7 @@ instance Icc_smooth_manifold (x y : ℝ) [Fact (x < y)] : either the left chart or the right chart, leaving 4 possibilities that we handle successively. -/ rcases he with (rfl | rfl) <;> rcases he' with (rfl | rfl) · -- `e = left chart`, `e' = left chart` - exact (mem_groupoid_of_pregroupoid.mpr (symm_trans_mem_contDiffGroupoid _ _ _)).1 + exact (mem_groupoid_of_pregroupoid.mpr (symm_trans_mem_contDiffGroupoid _)).1 · -- `e = left chart`, `e' = right chart` apply M.contDiffOn.congr rintro _ ⟨⟨hz₁, hz₂⟩, ⟨⟨z, hz₀⟩, rfl⟩⟩ @@ -347,7 +347,7 @@ instance Icc_smooth_manifold (x y : ℝ) [Fact (x < y)] : PiLp.neg_apply, update_same, max_eq_left, hz₀, hz₁.le, mfld_simps] abel ·-- `e = right chart`, `e' = right chart` - exact (mem_groupoid_of_pregroupoid.mpr (symm_trans_mem_contDiffGroupoid _ _ _)).1 + exact (mem_groupoid_of_pregroupoid.mpr (symm_trans_mem_contDiffGroupoid _)).1 /-! Register the manifold structure on `Icc 0 1`, and also its zero and one. -/ diff --git a/Mathlib/Geometry/Manifold/Instances/UnitsOfNormedAlgebra.lean b/Mathlib/Geometry/Manifold/Instances/UnitsOfNormedAlgebra.lean index afe2f7f2112cb..8ee99e793fb7c 100644 --- a/Mathlib/Geometry/Manifold/Instances/UnitsOfNormedAlgebra.lean +++ b/Mathlib/Geometry/Manifold/Instances/UnitsOfNormedAlgebra.lean @@ -44,12 +44,12 @@ theorem chartAt_source {a : Rˣ} : (chartAt R a).source = Set.univ := variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] [NormedAlgebra 𝕜 R] instance : SmoothManifoldWithCorners 𝓘(𝕜, R) Rˣ := - isOpenEmbedding_val.singleton_smoothManifoldWithCorners 𝓘(𝕜, R) + isOpenEmbedding_val.singleton_smoothManifoldWithCorners /-- For a complete normed ring `R`, the embedding of the units `Rˣ` into `R` is a smooth map between manifolds. -/ lemma contMDiff_val {m : ℕ∞} : ContMDiff 𝓘(𝕜, R) 𝓘(𝕜, R) m (val : Rˣ → R) := - contMDiff_isOpenEmbedding 𝓘(𝕜, R) Units.isOpenEmbedding_val + contMDiff_isOpenEmbedding Units.isOpenEmbedding_val /-- The units of a complete normed ring form a Lie group. -/ instance : LieGroup 𝓘(𝕜, R) Rˣ where diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index b16bd2c24cde9..624dbdbf85a0b 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -161,12 +161,12 @@ lemma IsIntegralCurveOn.hasDerivAt (hγ : IsIntegralCurveOn γ v s) {t : ℝ} (h have hsrc := extChartAt_source I (γ t₀) ▸ hsrc rw [hasDerivAt_iff_hasFDerivAt, ← hasMFDerivAt_iff_hasFDerivAt] apply (HasMFDerivAt.comp t - (hasMFDerivAt_extChartAt I hsrc) (hγ _ ht)).congr_mfderiv + (hasMFDerivAt_extChartAt (I := I) hsrc) (hγ _ ht)).congr_mfderiv rw [ContinuousLinearMap.ext_iff] intro a rw [ContinuousLinearMap.comp_apply, ContinuousLinearMap.smulRight_apply, map_smul, ← ContinuousLinearMap.one_apply (R₁ := ℝ) a, ← ContinuousLinearMap.smulRight_apply, - mfderiv_chartAt_eq_tangentCoordChange I hsrc] + mfderiv_chartAt_eq_tangentCoordChange hsrc] rfl variable [SmoothManifoldWithCorners I M] in @@ -174,17 +174,17 @@ lemma IsIntegralCurveAt.eventually_hasDerivAt (hγ : IsIntegralCurveAt γ v t₀ ∀ᶠ t in 𝓝 t₀, HasDerivAt ((extChartAt I (γ t₀)) ∘ γ) (tangentCoordChange I (γ t) (γ t₀) (γ t) (v (γ t))) t := by apply eventually_mem_nhds_iff.mpr - (hγ.continuousAt.preimage_mem_nhds (extChartAt_source_mem_nhds I _)) |>.and hγ |>.mono + (hγ.continuousAt.preimage_mem_nhds (extChartAt_source_mem_nhds (I := I) _)) |>.and hγ |>.mono rintro t ⟨ht1, ht2⟩ have hsrc := mem_of_mem_nhds ht1 rw [mem_preimage, extChartAt_source I (γ t₀)] at hsrc rw [hasDerivAt_iff_hasFDerivAt, ← hasMFDerivAt_iff_hasFDerivAt] - apply (HasMFDerivAt.comp t (hasMFDerivAt_extChartAt I hsrc) ht2).congr_mfderiv + apply (HasMFDerivAt.comp t (hasMFDerivAt_extChartAt (I := I) hsrc) ht2).congr_mfderiv rw [ContinuousLinearMap.ext_iff] intro a rw [ContinuousLinearMap.comp_apply, ContinuousLinearMap.smulRight_apply, map_smul, ← ContinuousLinearMap.one_apply (R₁ := ℝ) a, ← ContinuousLinearMap.smulRight_apply, - mfderiv_chartAt_eq_tangentCoordChange I hsrc] + mfderiv_chartAt_eq_tangentCoordChange hsrc] rfl /-! ### Translation lemmas -/ @@ -344,9 +344,9 @@ theorem exists_isIntegralCurveAt_of_contMDiffAt [CompleteSpace E] have hf3' := mem_of_mem_of_subset hf3 interior_subset have hft1 := mem_preimage.mp <| mem_of_mem_of_subset hf3' (extChartAt I x₀).target_subset_preimage_source - have hft2 := mem_extChartAt_source I xₜ + have hft2 := mem_extChartAt_source (I := I) xₜ -- express the derivative of the integral curve in the local chart - refine ⟨(continuousAt_extChartAt_symm'' _ hf3').comp h.continuousAt, + refine ⟨(continuousAt_extChartAt_symm'' hf3').comp h.continuousAt, HasDerivWithinAt.hasFDerivWithinAt ?_⟩ simp only [mfld_simps, hasDerivWithinAt_univ] show HasDerivAt ((extChartAt I xₜ ∘ (extChartAt I x₀).symm) ∘ f) (v xₜ) t @@ -368,7 +368,7 @@ lemma exists_isIntegralCurveAt_of_contMDiffAt_boundaryless [CompleteSpace E] [BoundarylessManifold I M] (hv : ContMDiffAt I I.tangent 1 (fun x ↦ (⟨x, v x⟩ : TangentBundle I M)) x₀) : ∃ γ : ℝ → M, γ t₀ = x₀ ∧ IsIntegralCurveAt γ v t₀ := - exists_isIntegralCurveAt_of_contMDiffAt t₀ hv (BoundarylessManifold.isInteriorPoint I) + exists_isIntegralCurveAt_of_contMDiffAt t₀ hv BoundarylessManifold.isInteriorPoint variable {t₀} @@ -394,7 +394,7 @@ theorem isIntegralCurveAt_eventuallyEq_of_contMDiffAt (hγt₀ : I.IsInteriorPoi -- internal lemmas to reduce code duplication have hsrc {g} (hg : IsIntegralCurveAt g v t₀) : ∀ᶠ t in 𝓝 t₀, g ⁻¹' (extChartAt I (g t₀)).source ∈ 𝓝 t := eventually_mem_nhds_iff.mpr <| - continuousAt_def.mp hg.continuousAt _ <| extChartAt_source_mem_nhds I (g t₀) + continuousAt_def.mp hg.continuousAt _ <| extChartAt_source_mem_nhds (g t₀) have hmem {g : ℝ → M} {t} (ht : g ⁻¹' (extChartAt I (g t₀)).source ∈ 𝓝 t) : g t ∈ (extChartAt I (g t₀)).source := mem_preimage.mp <| mem_of_mem_nhds ht have hdrv {g} (hg : IsIntegralCurveAt g v t₀) (h' : γ t₀ = g t₀) : ∀ᶠ t in 𝓝 t₀, @@ -407,7 +407,7 @@ theorem isIntegralCurveAt_eventuallyEq_of_contMDiffAt (hγt₀ : I.IsInteriorPoi apply ht2.congr_deriv congr <;> rw [Function.comp_apply, PartialEquiv.left_inv _ (hmem ht1)] - · apply ((continuousAt_extChartAt I (g t₀)).comp hg.continuousAt).preimage_mem_nhds + · apply ((continuousAt_extChartAt (g t₀)).comp hg.continuousAt).preimage_mem_nhds rw [Function.comp_apply, ← h'] exact hs have heq {g} (hg : IsIntegralCurveAt g v t₀) : @@ -425,7 +425,7 @@ theorem isIntegralCurveAt_eventuallyEq_of_contMDiffAt_boundaryless [Boundaryless (hv : ContMDiffAt I I.tangent 1 (fun x ↦ (⟨x, v x⟩ : TangentBundle I M)) (γ t₀)) (hγ : IsIntegralCurveAt γ v t₀) (hγ' : IsIntegralCurveAt γ' v t₀) (h : γ t₀ = γ' t₀) : γ =ᶠ[𝓝 t₀] γ' := - isIntegralCurveAt_eventuallyEq_of_contMDiffAt (BoundarylessManifold.isInteriorPoint I) hv hγ hγ' h + isIntegralCurveAt_eventuallyEq_of_contMDiffAt BoundarylessManifold.isInteriorPoint hv hγ hγ' h variable [T2Space M] {a b : ℝ} @@ -477,7 +477,7 @@ theorem isIntegralCurveOn_Ioo_eqOn_of_contMDiff_boundaryless [BoundarylessManifo (hγ : IsIntegralCurveOn γ v (Ioo a b)) (hγ' : IsIntegralCurveOn γ' v (Ioo a b)) (h : γ t₀ = γ' t₀) : EqOn γ γ' (Ioo a b) := isIntegralCurveOn_Ioo_eqOn_of_contMDiff - ht₀ (fun _ _ ↦ BoundarylessManifold.isInteriorPoint I) hv hγ hγ' h + ht₀ (fun _ _ ↦ BoundarylessManifold.isInteriorPoint) hv hγ hγ' h /-- Global integral curves are unique. @@ -500,7 +500,7 @@ theorem isIntegralCurve_eq_of_contMDiff (hγt : ∀ t, I.IsInteriorPoint (γ t)) theorem isIntegralCurve_Ioo_eq_of_contMDiff_boundaryless [BoundarylessManifold I M] (hv : ContMDiff I I.tangent 1 (fun x ↦ (⟨x, v x⟩ : TangentBundle I M))) (hγ : IsIntegralCurve γ v) (hγ' : IsIntegralCurve γ' v) (h : γ t₀ = γ' t₀) : γ = γ' := - isIntegralCurve_eq_of_contMDiff (fun _ ↦ BoundarylessManifold.isInteriorPoint I) hv hγ hγ' h + isIntegralCurve_eq_of_contMDiff (fun _ ↦ BoundarylessManifold.isInteriorPoint) hv hγ hγ' h /-- For a global integral curve `γ`, if it crosses itself at `a b : ℝ`, then it is periodic with period `a - b`. -/ diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index ae3c3a754f493..2fc03bd7d0a9a 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -50,14 +50,17 @@ open scoped Topology -- Let `M` be a manifold with corners over the pair `(E, H)`. variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] - {H : Type*} [TopologicalSpace H] (I : ModelWithCorners 𝕜 E H) + {H : Type*} [TopologicalSpace H] {I : ModelWithCorners 𝕜 E H} {M : Type*} [TopologicalSpace M] [ChartedSpace H M] namespace ModelWithCorners + +variable (I) in /-- `p ∈ M` is an interior point of a manifold `M` iff its image in the extended chart lies in the interior of the model space. -/ def IsInteriorPoint (x : M) := extChartAt I x x ∈ interior (range I) +variable (I) in /-- `p ∈ M` is a boundary point of a manifold `M` iff its image in the extended chart lies on the boundary of the model space. -/ def IsBoundaryPoint (x : M) := extChartAt I x x ∈ frontier (range I) @@ -68,8 +71,8 @@ protected def interior : Set M := { x : M | I.IsInteriorPoint x } lemma isInteriorPoint_iff {x : M} : I.IsInteriorPoint x ↔ extChartAt I x x ∈ interior (extChartAt I x).target := - ⟨fun h ↦ (chartAt H x).mem_interior_extend_target _ (mem_chart_target H x) h, - fun h ↦ PartialHomeomorph.interior_extend_target_subset_interior_range _ _ h⟩ + ⟨fun h ↦ (chartAt H x).mem_interior_extend_target (mem_chart_target H x) h, + fun h ↦ PartialHomeomorph.interior_extend_target_subset_interior_range _ h⟩ variable (M) in /-- The **boundary** of a manifold `M` is the set of its boundary points. -/ @@ -106,7 +109,6 @@ lemma compl_interior : (I.interior M)ᶜ = I.boundary M:= by lemma compl_boundary : (I.boundary M)ᶜ = I.interior M:= by rw [← compl_interior, compl_compl] -variable {I} in lemma _root_.range_mem_nhds_isInteriorPoint {x : M} (h : I.IsInteriorPoint x) : range I ∈ 𝓝 (extChartAt I x x) := by rw [mem_nhds_iff] @@ -126,11 +128,11 @@ variable [I.Boundaryless] /-- Boundaryless `ModelWithCorners` implies boundaryless manifold. -/ instance : BoundarylessManifold I M where isInteriorPoint' x := by - let r := ((chartAt H x).isOpen_extend_target I).interior_eq + let r := ((chartAt H x).isOpen_extend_target (I := I)).interior_eq have : extChartAt I x = (chartAt H x).extend I := rfl rw [← this] at r rw [isInteriorPoint_iff, r] - exact PartialEquiv.map_source _ (mem_extChartAt_source _ _) + exact PartialEquiv.map_source _ (mem_extChartAt_source _) end Boundaryless @@ -145,18 +147,18 @@ lemma _root_.BoundarylessManifold.isInteriorPoint {x : M} [BoundarylessManifold /-- If `I` is boundaryless, `M` has full interior. -/ lemma interior_eq_univ [BoundarylessManifold I M] : I.interior M = univ := - eq_univ_of_forall fun _ => BoundarylessManifold.isInteriorPoint I + eq_univ_of_forall fun _ => BoundarylessManifold.isInteriorPoint /-- Boundaryless manifolds have empty boundary. -/ lemma Boundaryless.boundary_eq_empty [BoundarylessManifold I M] : I.boundary M = ∅ := by rw [← I.compl_interior, I.interior_eq_univ, compl_empty_iff] instance [BoundarylessManifold I M] : IsEmpty (I.boundary M) := - isEmpty_coe_sort.mpr (Boundaryless.boundary_eq_empty I) + isEmpty_coe_sort.mpr Boundaryless.boundary_eq_empty /-- `M` is boundaryless iff its boundary is empty. -/ lemma Boundaryless.iff_boundary_eq_empty : I.boundary M = ∅ ↔ BoundarylessManifold I M := by - refine ⟨fun h ↦ { isInteriorPoint' := ?_ }, fun a ↦ boundary_eq_empty I⟩ + refine ⟨fun h ↦ { isInteriorPoint' := ?_ }, fun a ↦ boundary_eq_empty⟩ intro x show x ∈ I.interior M rw [← compl_interior, compl_empty_iff] at h @@ -172,11 +174,11 @@ end BoundarylessManifold /-! Interior and boundary of the product of two manifolds. -/ section prod -variable {I} +variable {E' : Type*} [NormedAddCommGroup E'] [NormedSpace 𝕜 E'] {H' : Type*} [TopologicalSpace H'] {N : Type*} [TopologicalSpace N] [ChartedSpace H' N] - (J : ModelWithCorners 𝕜 E' H') {x : M} {y : N} + {J : ModelWithCorners 𝕜 E' H'} {x : M} {y : N} /-- The interior of `M × N` is the product of the interiors of `M` and `N`. -/ lemma interior_prod : @@ -198,7 +200,7 @@ lemma interior_prod : lemma boundary_prod : (I.prod J).boundary (M × N) = Set.prod univ (J.boundary N) ∪ Set.prod (I.boundary M) univ := by let h := calc (I.prod J).boundary (M × N) - _ = ((I.prod J).interior (M × N))ᶜ := (I.prod J).compl_interior.symm + _ = ((I.prod J).interior (M × N))ᶜ := compl_interior.symm _ = ((I.interior M) ×ˢ (J.interior N))ᶜ := by rw [interior_prod] _ = (I.interior M)ᶜ ×ˢ univ ∪ univ ×ˢ (J.interior N)ᶜ := by rw [compl_prod_eq_union] rw [h, I.compl_interior, J.compl_interior, union_comm] @@ -207,14 +209,14 @@ lemma boundary_prod : /-- If `M` is boundaryless, `∂(M×N) = M × ∂N`. -/ lemma boundary_of_boundaryless_left [BoundarylessManifold I M] : (I.prod J).boundary (M × N) = Set.prod (univ : Set M) (J.boundary N) := by - rw [boundary_prod, Boundaryless.boundary_eq_empty I] + rw [boundary_prod, Boundaryless.boundary_eq_empty (I := I)] have : Set.prod (∅ : Set M) (univ : Set N) = ∅ := Set.empty_prod rw [this, union_empty] /-- If `N` is boundaryless, `∂(M×N) = ∂M × N`. -/ lemma boundary_of_boundaryless_right [BoundarylessManifold J N] : (I.prod J).boundary (M × N) = Set.prod (I.boundary M) (univ : Set N) := by - rw [boundary_prod, Boundaryless.boundary_eq_empty J] + rw [boundary_prod, Boundaryless.boundary_eq_empty (I := J)] have : Set.prod (univ : Set M) (∅ : Set N) = ∅ := Set.prod_empty rw [this, empty_union] diff --git a/Mathlib/Geometry/Manifold/MFDeriv/Atlas.lean b/Mathlib/Geometry/Manifold/MFDeriv/Atlas.lean index cbdbcb8afc391..f7b0ade18a085 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/Atlas.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/Atlas.lean @@ -32,15 +32,19 @@ open Bundle Set Topology variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] - (I : ModelWithCorners 𝕜 E H) {M : Type*} [TopologicalSpace M] [ChartedSpace H M] + {I : ModelWithCorners 𝕜 E H} {M : Type*} [TopologicalSpace M] [ChartedSpace H M] {E' : Type*} [NormedAddCommGroup E'] [NormedSpace 𝕜 E'] {H' : Type*} [TopologicalSpace H'] - (I' : ModelWithCorners 𝕜 E' H') {M' : Type*} [TopologicalSpace M'] [ChartedSpace H' M'] + {I' : ModelWithCorners 𝕜 E' H'} {M' : Type*} [TopologicalSpace M'] [ChartedSpace H' M'] {E'' : Type*} [NormedAddCommGroup E''] [NormedSpace 𝕜 E''] {H'' : Type*} [TopologicalSpace H''] - (I'' : ModelWithCorners 𝕜 E'' H'') {M'' : Type*} [TopologicalSpace M''] [ChartedSpace H'' M''] + {I'' : ModelWithCorners 𝕜 E'' H''} {M'' : Type*} [TopologicalSpace M''] [ChartedSpace H'' M''] section ModelWithCorners namespace ModelWithCorners +/- In general, the model with corner `I` is implicit in most theorems in differential geometry, but +this section is about `I` as a map, not as a parameter. Therefore, we make it explicit. -/ +variable (I) + /-! #### Model with corners -/ protected theorem hasMFDerivAt {x} : HasMFDerivAt I 𝓘(𝕜, E) I x (ContinuousLinearMap.id _ _) := @@ -98,7 +102,7 @@ theorem mdifferentiableAt_atlas (h : e ∈ atlas H M) {x : M} (hx : x ∈ e.sour · apply IsOpen.mem_nhds ((PartialHomeomorph.open_source _).preimage I.continuous_symm) mem.1 theorem mdifferentiableOn_atlas (h : e ∈ atlas H M) : MDifferentiableOn I I e e.source := - fun _x hx => (mdifferentiableAt_atlas I h hx).mdifferentiableWithinAt + fun _x hx => (mdifferentiableAt_atlas h hx).mdifferentiableWithinAt theorem mdifferentiableAt_atlas_symm (h : e ∈ atlas H M) {x : H} (hx : x ∈ e.target) : MDifferentiableAt I I e.symm x := by @@ -119,13 +123,13 @@ theorem mdifferentiableAt_atlas_symm (h : e ∈ atlas H M) {x : H} (hx : x ∈ e · apply IsOpen.mem_nhds ((PartialHomeomorph.open_source _).preimage I.continuous_symm) mem.1 theorem mdifferentiableOn_atlas_symm (h : e ∈ atlas H M) : MDifferentiableOn I I e.symm e.target := - fun _x hx => (mdifferentiableAt_atlas_symm I h hx).mdifferentiableWithinAt + fun _x hx => (mdifferentiableAt_atlas_symm h hx).mdifferentiableWithinAt theorem mdifferentiable_of_mem_atlas (h : e ∈ atlas H M) : e.MDifferentiable I I := - ⟨mdifferentiableOn_atlas I h, mdifferentiableOn_atlas_symm I h⟩ + ⟨mdifferentiableOn_atlas h, mdifferentiableOn_atlas_symm h⟩ theorem mdifferentiable_chart (x : M) : (chartAt H x).MDifferentiable I I := - mdifferentiable_of_mem_atlas _ (chart_mem_atlas _ _) + mdifferentiable_of_mem_atlas (chart_mem_atlas _ _) end Charts @@ -133,7 +137,6 @@ end Charts /-! ### Differentiable partial homeomorphisms -/ namespace PartialHomeomorph.MDifferentiable -variable {I I' I''} variable {e : PartialHomeomorph M M'} (he : e.MDifferentiable I I') {e' : PartialHomeomorph M' M''} include he @@ -151,7 +154,7 @@ theorem symm_comp_deriv {x : M} (hx : x ∈ e.source) : have : mfderiv I I (e.symm ∘ e) x = (mfderiv I' I e.symm (e x)).comp (mfderiv I I' e x) := mfderiv_comp x (he.mdifferentiableAt_symm (e.map_source hx)) (he.mdifferentiableAt hx) rw [← this] - have : mfderiv I I (_root_.id : M → M) x = ContinuousLinearMap.id _ _ := mfderiv_id I + have : mfderiv I I (_root_.id : M → M) x = ContinuousLinearMap.id _ _ := mfderiv_id rw [← this] apply Filter.EventuallyEq.mfderiv_eq have : e.source ∈ 𝓝 x := e.open_source.mem_nhds hx @@ -218,24 +221,22 @@ end PartialHomeomorph.MDifferentiable section extChartAt -variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedAddCommGroup E] - [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] (I : ModelWithCorners 𝕜 E H) {M : Type*} - [TopologicalSpace M] [ChartedSpace H M] [SmoothManifoldWithCorners I M] {s : Set M} {x y : M} +variable [SmoothManifoldWithCorners I M] {s : Set M} {x y : M} theorem hasMFDerivAt_extChartAt (h : y ∈ (chartAt H x).source) : HasMFDerivAt I 𝓘(𝕜, E) (extChartAt I x) y (mfderiv I I (chartAt H x) y : _) := - I.hasMFDerivAt.comp y ((mdifferentiable_chart I x).mdifferentiableAt h).hasMFDerivAt + I.hasMFDerivAt.comp y ((mdifferentiable_chart x).mdifferentiableAt h).hasMFDerivAt theorem hasMFDerivWithinAt_extChartAt (h : y ∈ (chartAt H x).source) : HasMFDerivWithinAt I 𝓘(𝕜, E) (extChartAt I x) s y (mfderiv I I (chartAt H x) y : _) := - (hasMFDerivAt_extChartAt I h).hasMFDerivWithinAt + (hasMFDerivAt_extChartAt h).hasMFDerivWithinAt theorem mdifferentiableAt_extChartAt (h : y ∈ (chartAt H x).source) : MDifferentiableAt I 𝓘(𝕜, E) (extChartAt I x) y := - (hasMFDerivAt_extChartAt I h).mdifferentiableAt + (hasMFDerivAt_extChartAt h).mdifferentiableAt theorem mdifferentiableOn_extChartAt : MDifferentiableOn I 𝓘(𝕜, E) (extChartAt I x) (chartAt H x).source := fun _y hy => - (hasMFDerivWithinAt_extChartAt I hy).mdifferentiableWithinAt + (hasMFDerivWithinAt_extChartAt hy).mdifferentiableWithinAt end extChartAt diff --git a/Mathlib/Geometry/Manifold/MFDeriv/Basic.lean b/Mathlib/Geometry/Manifold/MFDeriv/Basic.lean index b1835322451ff..883bc8a09c31f 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/Basic.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/Basic.lean @@ -123,12 +123,12 @@ theorem mdifferentiableWithinAt_univ : theorem mdifferentiableWithinAt_inter (ht : t ∈ 𝓝 x) : MDifferentiableWithinAt I I' f (s ∩ t) x ↔ MDifferentiableWithinAt I I' f s x := by rw [MDifferentiableWithinAt, MDifferentiableWithinAt, - (differentiableWithinAt_localInvariantProp I I').liftPropWithinAt_inter ht] + differentiableWithinAt_localInvariantProp.liftPropWithinAt_inter ht] theorem mdifferentiableWithinAt_inter' (ht : t ∈ 𝓝[s] x) : MDifferentiableWithinAt I I' f (s ∩ t) x ↔ MDifferentiableWithinAt I I' f s x := by rw [MDifferentiableWithinAt, MDifferentiableWithinAt, - (differentiableWithinAt_localInvariantProp I I').liftPropWithinAt_inter' ht] + differentiableWithinAt_localInvariantProp.liftPropWithinAt_inter' ht] theorem MDifferentiableAt.mdifferentiableWithinAt (h : MDifferentiableAt I I' f x) : MDifferentiableWithinAt I I' f s x := @@ -169,7 +169,7 @@ theorem ContMDiffWithinAt.mdifferentiableWithinAt (hf : ContMDiffWithinAt I I' n suffices h : MDifferentiableWithinAt I I' f (s ∩ f ⁻¹' (extChartAt I' (f x)).source) x by rwa [mdifferentiableWithinAt_inter'] at h apply hf.1.preimage_mem_nhdsWithin - exact extChartAt_source_mem_nhds I' (f x) + exact extChartAt_source_mem_nhds (f x) rw [mdifferentiableWithinAt_iff] exact ⟨hf.1.mono inter_subset_left, (hf.2.differentiableWithinAt hn).mono (by mfld_set_tac)⟩ @@ -255,8 +255,8 @@ theorem writtenInExtChartAt_comp (h : ContinuousWithinAt f s x) : 𝓝[(extChartAt I x).symm ⁻¹' s ∩ range I] (extChartAt I x) x := by apply @Filter.mem_of_superset _ _ (f ∘ (extChartAt I x).symm ⁻¹' (extChartAt I' (f x)).source) _ - (extChartAt_preimage_mem_nhdsWithin I - (h.preimage_mem_nhdsWithin (extChartAt_source_mem_nhds _ _))) + (extChartAt_preimage_mem_nhdsWithin + (h.preimage_mem_nhdsWithin (extChartAt_source_mem_nhds _))) mfld_set_tac variable {f' f₀' f₁' : TangentSpace I x →L[𝕜] TangentSpace I' (f x)} @@ -287,7 +287,7 @@ theorem mdifferentiableWithinAt_iff_of_mem_source {x' : M} {y : M'} ContinuousWithinAt f s x' ∧ DifferentiableWithinAt 𝕜 (extChartAt I' y ∘ f ∘ (extChartAt I x).symm) ((extChartAt I x).symm ⁻¹' s ∩ Set.range I) ((extChartAt I x) x') := - (differentiableWithinAt_localInvariantProp I I').liftPropWithinAt_indep_chart + differentiableWithinAt_localInvariantProp.liftPropWithinAt_indep_chart (StructureGroupoid.chart_mem_maximalAtlas _ x) hx (StructureGroupoid.chart_mem_maximalAtlas _ y) hy @@ -330,13 +330,13 @@ theorem hasMFDerivWithinAt_inter' (h : t ∈ 𝓝[s] x) : HasMFDerivWithinAt I I' f (s ∩ t) x f' ↔ HasMFDerivWithinAt I I' f s x f' := by rw [HasMFDerivWithinAt, HasMFDerivWithinAt, extChartAt_preimage_inter_eq, hasFDerivWithinAt_inter', continuousWithinAt_inter' h] - exact extChartAt_preimage_mem_nhdsWithin I h + exact extChartAt_preimage_mem_nhdsWithin h theorem hasMFDerivWithinAt_inter (h : t ∈ 𝓝 x) : HasMFDerivWithinAt I I' f (s ∩ t) x f' ↔ HasMFDerivWithinAt I I' f s x f' := by rw [HasMFDerivWithinAt, HasMFDerivWithinAt, extChartAt_preimage_inter_eq, hasFDerivWithinAt_inter, continuousWithinAt_inter h] - exact extChartAt_preimage_mem_nhds I h + exact extChartAt_preimage_mem_nhds h theorem HasMFDerivWithinAt.union (hs : HasMFDerivWithinAt I I' f s x f') (ht : HasMFDerivWithinAt I I' f t x f') : HasMFDerivWithinAt I I' f (s ∪ t) x f' := by @@ -403,7 +403,7 @@ theorem mfderivWithin_univ : mfderivWithin I I' f univ = mfderiv I I' f := by theorem mfderivWithin_inter (ht : t ∈ 𝓝 x) : mfderivWithin I I' f (s ∩ t) x = mfderivWithin I I' f s x := by rw [mfderivWithin, mfderivWithin, extChartAt_preimage_inter_eq, mdifferentiableWithinAt_inter ht, - fderivWithin_inter (extChartAt_preimage_mem_nhds I ht)] + fderivWithin_inter (extChartAt_preimage_mem_nhds ht)] theorem mfderivWithin_of_mem_nhds (h : s ∈ 𝓝 x) : mfderivWithin I I' f s x = mfderiv I I' f x := by rw [← mfderivWithin_univ, ← univ_inter s, mfderivWithin_inter h] @@ -479,7 +479,7 @@ theorem HasMFDerivWithinAt.congr_of_eventuallyEq (h : HasMFDerivWithinAt I I' f · have : (extChartAt I x).symm ⁻¹' {y | f₁ y = f y} ∈ 𝓝[(extChartAt I x).symm ⁻¹' s ∩ range I] (extChartAt I x) x := - extChartAt_preimage_mem_nhdsWithin I h₁ + extChartAt_preimage_mem_nhdsWithin h₁ apply Filter.mem_of_superset this fun y => _ simp (config := { contextual := true }) only [hx, mfld_simps] · simp only [hx, mfld_simps] @@ -586,8 +586,8 @@ theorem HasMFDerivWithinAt.comp (hg : HasMFDerivWithinAt I' I'' g u (f x) g') have : (extChartAt I x).symm ⁻¹' (f ⁻¹' (extChartAt I' (f x)).source) ∈ 𝓝[(extChartAt I x).symm ⁻¹' s ∩ range I] (extChartAt I x) x := - extChartAt_preimage_mem_nhdsWithin I - (hf.1.preimage_mem_nhdsWithin (extChartAt_source_mem_nhds _ _)) + extChartAt_preimage_mem_nhdsWithin + (hf.1.preimage_mem_nhdsWithin (extChartAt_source_mem_nhds _)) unfold HasMFDerivWithinAt at * rw [← hasFDerivWithinAt_inter' this, ← extChartAt_preimage_inter_eq] at hf ⊢ have : writtenInExtChartAt I I' x f ((extChartAt I x) x) = (extChartAt I' (f x)) (f x) := by diff --git a/Mathlib/Geometry/Manifold/MFDeriv/Defs.lean b/Mathlib/Geometry/Manifold/MFDeriv/Defs.lean index 88e42b72004c0..6d30a6a3b1c43 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/Defs.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/Defs.lean @@ -116,11 +116,12 @@ We use the names `MDifferentiable` and `mfderiv`, where the prefix letter `m` me -/ variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedAddCommGroup E] - [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] (I : ModelWithCorners 𝕜 E H) {M : Type*} + [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] {I : ModelWithCorners 𝕜 E H} {M : Type*} [TopologicalSpace M] [ChartedSpace H M] {E' : Type*} [NormedAddCommGroup E'] [NormedSpace 𝕜 E'] - {H' : Type*} [TopologicalSpace H'] (I' : ModelWithCorners 𝕜 E' H') {M' : Type*} + {H' : Type*} [TopologicalSpace H'] {I' : ModelWithCorners 𝕜 E' H'} {M' : Type*} [TopologicalSpace M'] [ChartedSpace H' M'] +variable (I I') in /-- Property in the model space of a model with corners of being differentiable within at set at a point, when read in the model vector space. This property will be lifted to manifolds to define differentiable functions between manifolds. -/ @@ -177,15 +178,18 @@ theorem differentiableWithinAt_localInvariantProp : @[deprecated (since := "2024-10-10")] alias differentiable_within_at_localInvariantProp := differentiableWithinAt_localInvariantProp +variable (I) in /-- Predicate ensuring that, at a point and within a set, a function can have at most one derivative. This is expressed using the preferred chart at the considered point. -/ def UniqueMDiffWithinAt (s : Set M) (x : M) := UniqueDiffWithinAt 𝕜 ((extChartAt I x).symm ⁻¹' s ∩ range I) ((extChartAt I x) x) +variable (I) in /-- Predicate ensuring that, at all points of a set, a function can have at most one derivative. -/ def UniqueMDiffOn (s : Set M) := ∀ x ∈ s, UniqueMDiffWithinAt I s x +variable (I I') in /-- `MDifferentiableWithinAt I I' f s x` indicates that the function `f` between manifolds has a derivative at the point `x` within the set `s`. This is a generalization of `DifferentiableWithinAt` to manifolds. @@ -206,19 +210,18 @@ theorem mdifferentiableWithinAt_iff' (f : M → M') (s : Set M) (x : M) : @[deprecated (since := "2024-04-30")] alias mdifferentiableWithinAt_iff_liftPropWithinAt := mdifferentiableWithinAt_iff' -variable {I I'} in theorem MDifferentiableWithinAt.continuousWithinAt {f : M → M'} {s : Set M} {x : M} (hf : MDifferentiableWithinAt I I' f s x) : ContinuousWithinAt f s x := mdifferentiableWithinAt_iff' .. |>.1 hf |>.1 -variable {I I'} in theorem MDifferentiableWithinAt.differentiableWithinAt_writtenInExtChartAt {f : M → M'} {s : Set M} {x : M} (hf : MDifferentiableWithinAt I I' f s x) : DifferentiableWithinAt 𝕜 (writtenInExtChartAt I I' x f) ((extChartAt I x).symm ⁻¹' s ∩ range I) ((extChartAt I x) x) := mdifferentiableWithinAt_iff' .. |>.1 hf |>.2 +variable (I I') in /-- `MDifferentiableAt I I' f x` indicates that the function `f` between manifolds has a derivative at the point `x`. This is a generalization of `DifferentiableAt` to manifolds. @@ -242,35 +245,35 @@ theorem mdifferentiableAt_iff (f : M → M') (x : M) : @[deprecated (since := "2024-04-30")] alias mdifferentiableAt_iff_liftPropAt := mdifferentiableAt_iff -variable {I I'} in theorem MDifferentiableAt.continuousAt {f : M → M'} {x : M} (hf : MDifferentiableAt I I' f x) : ContinuousAt f x := mdifferentiableAt_iff .. |>.1 hf |>.1 -variable {I I'} in theorem MDifferentiableAt.differentiableWithinAt_writtenInExtChartAt {f : M → M'} {x : M} (hf : MDifferentiableAt I I' f x) : DifferentiableWithinAt 𝕜 (writtenInExtChartAt I I' x f) (range I) ((extChartAt I x) x) := mdifferentiableAt_iff .. |>.1 hf |>.2 +variable (I I') in /-- `MDifferentiableOn I I' f s` indicates that the function `f` between manifolds has a derivative within `s` at all points of `s`. This is a generalization of `DifferentiableOn` to manifolds. -/ def MDifferentiableOn (f : M → M') (s : Set M) := ∀ x ∈ s, MDifferentiableWithinAt I I' f s x +variable (I I') in /-- `MDifferentiable I I' f` indicates that the function `f` between manifolds has a derivative everywhere. This is a generalization of `Differentiable` to manifolds. -/ def MDifferentiable (f : M → M') := ∀ x, MDifferentiableAt I I' f x +variable (I I') in /-- Prop registering if a partial homeomorphism is a local diffeomorphism on its source -/ def PartialHomeomorph.MDifferentiable (f : PartialHomeomorph M M') := MDifferentiableOn I I' f f.source ∧ MDifferentiableOn I' I f.symm f.target -variable [SmoothManifoldWithCorners I M] [SmoothManifoldWithCorners I' M'] - +variable (I I') in /-- `HasMFDerivWithinAt I I' f s x f'` indicates that the function `f` between manifolds has, at the point `x` and within the set `s`, the derivative `f'`. Here, `f'` is a continuous linear map from the tangent space at `x` to the tangent space at `f x`. @@ -288,6 +291,7 @@ def HasMFDerivWithinAt (f : M → M') (s : Set M) (x : M) HasFDerivWithinAt (writtenInExtChartAt I I' x f : E → E') f' ((extChartAt I x).symm ⁻¹' s ∩ range I) ((extChartAt I x) x) +variable (I I') in /-- `HasMFDerivAt I I' f x f'` indicates that the function `f` between manifolds has, at the point `x`, the derivative `f'`. Here, `f'` is a continuous linear map from the tangent space at `x` to the tangent space at `f x`. @@ -301,6 +305,7 @@ def HasMFDerivAt (f : M → M') (x : M) (f' : TangentSpace I x →L[𝕜] Tangen HasFDerivWithinAt (writtenInExtChartAt I I' x f : E → E') f' (range I) ((extChartAt I x) x) open Classical in +variable (I I') in /-- Let `f` be a function between two smooth manifolds. Then `mfderivWithin I I' f s x` is the derivative of `f` at `x` within `s`, as a continuous linear map from the tangent space at `x` to the tangent space at `f x`. -/ @@ -312,6 +317,7 @@ def mfderivWithin (f : M → M') (s : Set M) (x : M) : TangentSpace I x →L[ else 0 open Classical in +variable (I I') in /-- Let `f` be a function between two smooth manifolds. Then `mfderiv I I' f x` is the derivative of `f` at `x`, as a continuous linear map from the tangent space at `x` to the tangent space at `f x`. -/ @@ -320,10 +326,12 @@ def mfderiv (f : M → M') (x : M) : TangentSpace I x →L[𝕜] TangentSpace I' (fderivWithin 𝕜 (writtenInExtChartAt I I' x f : E → E') (range I) ((extChartAt I x) x) : _) else 0 +variable (I I') in /-- The derivative within a set, as a map between the tangent bundles -/ def tangentMapWithin (f : M → M') (s : Set M) : TangentBundle I M → TangentBundle I' M' := fun p => ⟨f p.1, (mfderivWithin I I' f s p.1 : TangentSpace I p.1 → TangentSpace I' (f p.1)) p.2⟩ +variable (I I') in /-- The derivative, as a map between the tangent bundles -/ def tangentMap (f : M → M') : TangentBundle I M → TangentBundle I' M' := fun p => ⟨f p.1, (mfderiv I I' f p.1 : TangentSpace I p.1 → TangentSpace I' (f p.1)) p.2⟩ diff --git a/Mathlib/Geometry/Manifold/MFDeriv/SpecificFunctions.lean b/Mathlib/Geometry/Manifold/MFDeriv/SpecificFunctions.lean index e80b20105886e..2bbe312b6c383 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/SpecificFunctions.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/SpecificFunctions.lean @@ -27,12 +27,12 @@ section SpecificFunctions /-! ### Differentiability of specific functions -/ variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedAddCommGroup E] - [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] (I : ModelWithCorners 𝕜 E H) {M : Type*} + [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] {I : ModelWithCorners 𝕜 E H} {M : Type*} [TopologicalSpace M] [ChartedSpace H M] {E' : Type*} [NormedAddCommGroup E'] [NormedSpace 𝕜 E'] {H' : Type*} [TopologicalSpace H'] - (I' : ModelWithCorners 𝕜 E' H') {M' : Type*} [TopologicalSpace M'] [ChartedSpace H' M'] + {I' : ModelWithCorners 𝕜 E' H'} {M' : Type*} [TopologicalSpace M'] [ChartedSpace H' M'] {E'' : Type*} [NormedAddCommGroup E''] [NormedSpace 𝕜 E''] - {H'' : Type*} [TopologicalSpace H''] (I'' : ModelWithCorners 𝕜 E'' H'') {M'' : Type*} + {H'' : Type*} [TopologicalSpace H''] {I'' : ModelWithCorners 𝕜 E'' H''} {M'' : Type*} [TopologicalSpace M''] [ChartedSpace H'' M''] namespace ContinuousLinearMap @@ -107,34 +107,34 @@ theorem hasMFDerivAt_id (x : M) : HasMFDerivAt I I (@id M) x (ContinuousLinearMap.id 𝕜 (TangentSpace I x)) := by refine ⟨continuousAt_id, ?_⟩ have : ∀ᶠ y in 𝓝[range I] (extChartAt I x) x, (extChartAt I x ∘ (extChartAt I x).symm) y = y := by - apply Filter.mem_of_superset (extChartAt_target_mem_nhdsWithin I x) + apply Filter.mem_of_superset (extChartAt_target_mem_nhdsWithin x) mfld_set_tac apply HasFDerivWithinAt.congr_of_eventuallyEq (hasFDerivWithinAt_id _ _) this simp only [mfld_simps] theorem hasMFDerivWithinAt_id (s : Set M) (x : M) : HasMFDerivWithinAt I I (@id M) s x (ContinuousLinearMap.id 𝕜 (TangentSpace I x)) := - (hasMFDerivAt_id I x).hasMFDerivWithinAt + (hasMFDerivAt_id x).hasMFDerivWithinAt theorem mdifferentiableAt_id : MDifferentiableAt I I (@id M) x := - (hasMFDerivAt_id I x).mdifferentiableAt + (hasMFDerivAt_id x).mdifferentiableAt theorem mdifferentiableWithinAt_id : MDifferentiableWithinAt I I (@id M) s x := - (mdifferentiableAt_id I).mdifferentiableWithinAt + mdifferentiableAt_id.mdifferentiableWithinAt -theorem mdifferentiable_id : MDifferentiable I I (@id M) := fun _ => mdifferentiableAt_id I +theorem mdifferentiable_id : MDifferentiable I I (@id M) := fun _ => mdifferentiableAt_id theorem mdifferentiableOn_id : MDifferentiableOn I I (@id M) s := - (mdifferentiable_id I).mdifferentiableOn + mdifferentiable_id.mdifferentiableOn @[simp, mfld_simps] theorem mfderiv_id : mfderiv I I (@id M) x = ContinuousLinearMap.id 𝕜 (TangentSpace I x) := - HasMFDerivAt.mfderiv (hasMFDerivAt_id I x) + HasMFDerivAt.mfderiv (hasMFDerivAt_id x) theorem mfderivWithin_id (hxs : UniqueMDiffWithinAt I s x) : mfderivWithin I I (@id M) s x = ContinuousLinearMap.id 𝕜 (TangentSpace I x) := by - rw [MDifferentiable.mfderivWithin (mdifferentiableAt_id I) hxs] - exact mfderiv_id I + rw [MDifferentiable.mfderivWithin mdifferentiableAt_id hxs] + exact mfderiv_id @[simp, mfld_simps] theorem tangentMap_id : tangentMap I I (id : M → M) = id := by ext1 ⟨x, v⟩; simp [tangentMap] @@ -162,28 +162,28 @@ theorem hasMFDerivAt_const (c : M') (x : M) : theorem hasMFDerivWithinAt_const (c : M') (s : Set M) (x : M) : HasMFDerivWithinAt I I' (fun _ : M => c) s x (0 : TangentSpace I x →L[𝕜] TangentSpace I' c) := - (hasMFDerivAt_const I I' c x).hasMFDerivWithinAt + (hasMFDerivAt_const c x).hasMFDerivWithinAt theorem mdifferentiableAt_const : MDifferentiableAt I I' (fun _ : M => c) x := - (hasMFDerivAt_const I I' c x).mdifferentiableAt + (hasMFDerivAt_const c x).mdifferentiableAt theorem mdifferentiableWithinAt_const : MDifferentiableWithinAt I I' (fun _ : M => c) s x := - (mdifferentiableAt_const I I').mdifferentiableWithinAt + mdifferentiableAt_const.mdifferentiableWithinAt theorem mdifferentiable_const : MDifferentiable I I' fun _ : M => c := fun _ => - mdifferentiableAt_const I I' + mdifferentiableAt_const theorem mdifferentiableOn_const : MDifferentiableOn I I' (fun _ : M => c) s := - (mdifferentiable_const I I').mdifferentiableOn + mdifferentiable_const.mdifferentiableOn @[simp, mfld_simps] theorem mfderiv_const : mfderiv I I' (fun _ : M => c) x = (0 : TangentSpace I x →L[𝕜] TangentSpace I' c) := - HasMFDerivAt.mfderiv (hasMFDerivAt_const I I' c x) + HasMFDerivAt.mfderiv (hasMFDerivAt_const c x) theorem mfderivWithin_const (hxs : UniqueMDiffWithinAt I s x) : mfderivWithin I I' (fun _ : M => c) s x = (0 : TangentSpace I x →L[𝕜] TangentSpace I' c) := - (hasMFDerivWithinAt_const _ _ _ _ _).mfderivWithin hxs + (hasMFDerivWithinAt_const _ _ _).mfderivWithin hxs end Const @@ -202,42 +202,42 @@ theorem hasMFDerivAt_fst (x : M × M') : apply Filter.mem_of_superset (extChartAt_target_mem_nhdsWithin (I.prod I') x) mfld_set_tac -/ - filter_upwards [extChartAt_target_mem_nhdsWithin (I.prod I') x] with y hy + filter_upwards [extChartAt_target_mem_nhdsWithin x] with y hy rw [extChartAt_prod] at hy exact (extChartAt I x.1).right_inv hy.1 apply HasFDerivWithinAt.congr_of_eventuallyEq hasFDerivWithinAt_fst this -- Porting note: next line was `simp only [mfld_simps]` - exact (extChartAt I x.1).right_inv <| (extChartAt I x.1).map_source (mem_extChartAt_source _ _) + exact (extChartAt I x.1).right_inv <| (extChartAt I x.1).map_source (mem_extChartAt_source _) theorem hasMFDerivWithinAt_fst (s : Set (M × M')) (x : M × M') : HasMFDerivWithinAt (I.prod I') I Prod.fst s x (ContinuousLinearMap.fst 𝕜 (TangentSpace I x.1) (TangentSpace I' x.2)) := - (hasMFDerivAt_fst I I' x).hasMFDerivWithinAt + (hasMFDerivAt_fst x).hasMFDerivWithinAt theorem mdifferentiableAt_fst {x : M × M'} : MDifferentiableAt (I.prod I') I Prod.fst x := - (hasMFDerivAt_fst I I' x).mdifferentiableAt + (hasMFDerivAt_fst x).mdifferentiableAt theorem mdifferentiableWithinAt_fst {s : Set (M × M')} {x : M × M'} : MDifferentiableWithinAt (I.prod I') I Prod.fst s x := - (mdifferentiableAt_fst I I').mdifferentiableWithinAt + mdifferentiableAt_fst.mdifferentiableWithinAt theorem mdifferentiable_fst : MDifferentiable (I.prod I') I (Prod.fst : M × M' → M) := fun _ => - mdifferentiableAt_fst I I' + mdifferentiableAt_fst theorem mdifferentiableOn_fst {s : Set (M × M')} : MDifferentiableOn (I.prod I') I Prod.fst s := - (mdifferentiable_fst I I').mdifferentiableOn + mdifferentiable_fst.mdifferentiableOn @[simp, mfld_simps] theorem mfderiv_fst {x : M × M'} : mfderiv (I.prod I') I Prod.fst x = ContinuousLinearMap.fst 𝕜 (TangentSpace I x.1) (TangentSpace I' x.2) := - (hasMFDerivAt_fst I I' x).mfderiv + (hasMFDerivAt_fst x).mfderiv theorem mfderivWithin_fst {s : Set (M × M')} {x : M × M'} (hxs : UniqueMDiffWithinAt (I.prod I') s x) : mfderivWithin (I.prod I') I Prod.fst s x = ContinuousLinearMap.fst 𝕜 (TangentSpace I x.1) (TangentSpace I' x.2) := by - rw [MDifferentiable.mfderivWithin (mdifferentiableAt_fst I I') hxs]; exact mfderiv_fst I I' + rw [MDifferentiable.mfderivWithin mdifferentiableAt_fst hxs]; exact mfderiv_fst @[simp, mfld_simps] theorem tangentMap_prod_fst {p : TangentBundle (I.prod I') (M × M')} : @@ -264,42 +264,42 @@ theorem hasMFDerivAt_snd (x : M × M') : apply Filter.mem_of_superset (extChartAt_target_mem_nhdsWithin (I.prod I') x) mfld_set_tac -/ - filter_upwards [extChartAt_target_mem_nhdsWithin (I.prod I') x] with y hy + filter_upwards [extChartAt_target_mem_nhdsWithin x] with y hy rw [extChartAt_prod] at hy exact (extChartAt I' x.2).right_inv hy.2 apply HasFDerivWithinAt.congr_of_eventuallyEq hasFDerivWithinAt_snd this -- Porting note: the next line was `simp only [mfld_simps]` - exact (extChartAt I' x.2).right_inv <| (extChartAt I' x.2).map_source (mem_extChartAt_source _ _) + exact (extChartAt I' x.2).right_inv <| (extChartAt I' x.2).map_source (mem_extChartAt_source _) theorem hasMFDerivWithinAt_snd (s : Set (M × M')) (x : M × M') : HasMFDerivWithinAt (I.prod I') I' Prod.snd s x (ContinuousLinearMap.snd 𝕜 (TangentSpace I x.1) (TangentSpace I' x.2)) := - (hasMFDerivAt_snd I I' x).hasMFDerivWithinAt + (hasMFDerivAt_snd x).hasMFDerivWithinAt theorem mdifferentiableAt_snd {x : M × M'} : MDifferentiableAt (I.prod I') I' Prod.snd x := - (hasMFDerivAt_snd I I' x).mdifferentiableAt + (hasMFDerivAt_snd x).mdifferentiableAt theorem mdifferentiableWithinAt_snd {s : Set (M × M')} {x : M × M'} : MDifferentiableWithinAt (I.prod I') I' Prod.snd s x := - (mdifferentiableAt_snd I I').mdifferentiableWithinAt + mdifferentiableAt_snd.mdifferentiableWithinAt theorem mdifferentiable_snd : MDifferentiable (I.prod I') I' (Prod.snd : M × M' → M') := fun _ => - mdifferentiableAt_snd I I' + mdifferentiableAt_snd theorem mdifferentiableOn_snd {s : Set (M × M')} : MDifferentiableOn (I.prod I') I' Prod.snd s := - (mdifferentiable_snd I I').mdifferentiableOn + mdifferentiable_snd.mdifferentiableOn @[simp, mfld_simps] theorem mfderiv_snd {x : M × M'} : mfderiv (I.prod I') I' Prod.snd x = ContinuousLinearMap.snd 𝕜 (TangentSpace I x.1) (TangentSpace I' x.2) := - (hasMFDerivAt_snd I I' x).mfderiv + (hasMFDerivAt_snd x).mfderiv theorem mfderivWithin_snd {s : Set (M × M')} {x : M × M'} (hxs : UniqueMDiffWithinAt (I.prod I') s x) : mfderivWithin (I.prod I') I' Prod.snd s x = ContinuousLinearMap.snd 𝕜 (TangentSpace I x.1) (TangentSpace I' x.2) := by - rw [MDifferentiable.mfderivWithin (mdifferentiableAt_snd I I') hxs]; exact mfderiv_snd I I' + rw [MDifferentiable.mfderivWithin mdifferentiableAt_snd hxs]; exact mfderiv_snd @[simp, mfld_simps] theorem tangentMap_prod_snd {p : TangentBundle (I.prod I') (M × M')} : @@ -315,8 +315,6 @@ theorem tangentMapWithin_prod_snd {s : Set (M × M')} {p : TangentBundle (I.prod · rcases p with ⟨⟩; rfl · exact hs -variable {I I' I''} - theorem MDifferentiableAt.mfderiv_prod {f : M → M'} {g : M → M''} {x : M} (hf : MDifferentiableAt I I' f x) (hg : MDifferentiableAt I I'' g x) : mfderiv I (I'.prod I'') (fun x => (f x, g x)) x = @@ -326,18 +324,16 @@ theorem MDifferentiableAt.mfderiv_prod {f : M → M'} {g : M → M''} {x : M} exact hf.differentiableWithinAt_writtenInExtChartAt.fderivWithin_prod hg.differentiableWithinAt_writtenInExtChartAt (I.uniqueDiffOn _ (mem_range_self _)) -variable (I I' I'') - theorem mfderiv_prod_left {x₀ : M} {y₀ : M'} : mfderiv I (I.prod I') (fun x => (x, y₀)) x₀ = ContinuousLinearMap.inl 𝕜 (TangentSpace I x₀) (TangentSpace I' y₀) := by - refine ((mdifferentiableAt_id I).mfderiv_prod (mdifferentiableAt_const I I')).trans ?_ + refine (mdifferentiableAt_id.mfderiv_prod mdifferentiableAt_const).trans ?_ rw [mfderiv_id, mfderiv_const, ContinuousLinearMap.inl] theorem mfderiv_prod_right {x₀ : M} {y₀ : M'} : mfderiv I' (I.prod I') (fun y => (x₀, y)) y₀ = ContinuousLinearMap.inr 𝕜 (TangentSpace I x₀) (TangentSpace I' y₀) := by - refine ((mdifferentiableAt_const I' I).mfderiv_prod (mdifferentiableAt_id I')).trans ?_ + refine (mdifferentiableAt_const.mfderiv_prod mdifferentiableAt_id).trans ?_ rw [mfderiv_id, mfderiv_const, ContinuousLinearMap.inr] /-- The total derivative of a function in two variables is the sum of the partial derivatives. @@ -350,12 +346,11 @@ theorem mfderiv_prod_eq_add {f : M × M' → M''} {p : M × M'} mfderiv (I.prod I') I'' (fun z : M × M' => f (z.1, p.2)) p + mfderiv (I.prod I') I'' (fun z : M × M' => f (p.1, z.2)) p := by dsimp only - erw [mfderiv_comp_of_eq hf ((mdifferentiableAt_fst I I').prod_mk (mdifferentiableAt_const _ _)) - rfl, - mfderiv_comp_of_eq hf ((mdifferentiableAt_const _ _).prod_mk (mdifferentiableAt_snd I I')) rfl, + erw [mfderiv_comp_of_eq hf (mdifferentiableAt_fst.prod_mk mdifferentiableAt_const) rfl, + mfderiv_comp_of_eq hf (mdifferentiableAt_const.prod_mk mdifferentiableAt_snd) rfl, ← ContinuousLinearMap.comp_add, - (mdifferentiableAt_fst I I').mfderiv_prod (mdifferentiableAt_const (I.prod I') I'), - (mdifferentiableAt_const (I.prod I') I).mfderiv_prod (mdifferentiableAt_snd I I'), mfderiv_fst, + mdifferentiableAt_fst.mfderiv_prod mdifferentiableAt_const, + mdifferentiableAt_const.mfderiv_prod mdifferentiableAt_snd, mfderiv_fst, mfderiv_snd, mfderiv_const, mfderiv_const] symm convert ContinuousLinearMap.comp_id <| mfderiv (.prod I I') I'' f (p.1, p.2) @@ -374,7 +369,6 @@ canonical, but in this case (the tangent space of a vector space) it is canonica section Group -variable {I} variable {z : M} {f g : M → E'} {f' g' : TangentSpace I z →L[𝕜] E'} theorem HasMFDerivAt.add (hf : HasMFDerivAt I 𝓘(𝕜, E') f z f') @@ -460,7 +454,6 @@ end Group section AlgebraOverRing -variable {I} variable {z : M} {F' : Type*} [NormedRing F'] [NormedAlgebra 𝕜 F'] {p q : M → F'} {p' q' : TangentSpace I z →L[𝕜] F'} @@ -495,7 +488,6 @@ end AlgebraOverRing section AlgebraOverCommRing -variable {I} variable {z : M} {F' : Type*} [NormedCommRing F'] [NormedAlgebra 𝕜 F'] {p q : M → F'} {p' q' : TangentSpace I z →L[𝕜] F'} diff --git a/Mathlib/Geometry/Manifold/MFDeriv/Tangent.lean b/Mathlib/Geometry/Manifold/MFDeriv/Tangent.lean index 633857d4ac0f7..4814f3ebcafdf 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/Tangent.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/Tangent.lean @@ -19,7 +19,7 @@ open Bundle variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] - (I : ModelWithCorners 𝕜 E H) {M : Type*} [TopologicalSpace M] [ChartedSpace H M] + {I : ModelWithCorners 𝕜 E H} {M : Type*} [TopologicalSpace M] [ChartedSpace H M] [SmoothManifoldWithCorners I M] /-- The derivative of the chart at a base point is the chart of the tangent bundle, composed with @@ -31,7 +31,7 @@ theorem tangentMap_chart {p q : TangentBundle I M} (h : q.1 ∈ (chartAt H p.1). dsimp [tangentMap] rw [MDifferentiableAt.mfderiv] · rfl - · exact mdifferentiableAt_atlas _ (chart_mem_atlas _ _) h + · exact mdifferentiableAt_atlas (chart_mem_atlas _ _) h /-- The derivative of the inverse of the chart at a base point is the inverse of the chart of the tangent bundle, composed with the identification between the tangent bundle of the model space and @@ -41,7 +41,7 @@ theorem tangentMap_chart_symm {p : TangentBundle I M} {q : TangentBundle I H} tangentMap I I (chartAt H p.1).symm q = (chartAt (ModelProd H E) p).symm (TotalSpace.toProd H E q) := by dsimp only [tangentMap] - rw [MDifferentiableAt.mfderiv (mdifferentiableAt_atlas_symm _ (chart_mem_atlas _ _) h)] + rw [MDifferentiableAt.mfderiv (mdifferentiableAt_atlas_symm (chart_mem_atlas _ _) h)] simp only [ContinuousLinearMap.coe_coe, TangentBundle.chartAt, h, tangentBundleCore, mfld_simps, (· ∘ ·)] -- `simp` fails to apply `PartialEquiv.prod_symm` with `ModelProd` @@ -50,7 +50,7 @@ theorem tangentMap_chart_symm {p : TangentBundle I M} {q : TangentBundle I H} lemma mfderiv_chartAt_eq_tangentCoordChange {x y : M} (hsrc : x ∈ (chartAt H y).source) : mfderiv I I (chartAt H y) x = tangentCoordChange I x y x := by - have := mdifferentiableAt_atlas I (ChartedSpace.chart_mem_atlas _) hsrc + have := mdifferentiableAt_atlas (I := I) (ChartedSpace.chart_mem_atlas _) hsrc simp [mfderiv, if_pos this, Function.comp_assoc] /-- The preimage under the projection from the tangent bundle of a set with unique differential in diff --git a/Mathlib/Geometry/Manifold/MFDeriv/UniqueDifferential.lean b/Mathlib/Geometry/Manifold/MFDeriv/UniqueDifferential.lean index 076089056c123..4facc7f820910 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/UniqueDifferential.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/UniqueDifferential.lean @@ -44,7 +44,7 @@ theorem UniqueMDiffWithinAt.image_denseRange (hs : UniqueMDiffWithinAt I s x) {f : M → M'} {f' : E →L[𝕜] E'} (hf : HasMFDerivWithinAt I I' f s x f') (hd : DenseRange f') : UniqueMDiffWithinAt I' (f '' s) (f x) := by /- Rewrite in coordinates, apply `HasFDerivWithinAt.uniqueDiffWithinAt`. -/ - have := hs.inter' <| hf.1 (extChartAt_source_mem_nhds I' (f x)) + have := hs.inter' <| hf.1 (extChartAt_source_mem_nhds (I := I') (f x)) refine (((hf.2.mono ?sub1).uniqueDiffWithinAt this hd).mono ?sub2).congr_pt ?pt case pt => simp only [mfld_simps] case sub1 => mfld_set_tac @@ -92,8 +92,8 @@ theorem UniqueMDiffOn.uniqueDiffOn_target_inter (hs : UniqueMDiffOn I s) (x : M) apply UniqueMDiffOn.uniqueDiffOn rw [← PartialEquiv.image_source_inter_eq', inter_comm, extChartAt_source] exact (hs.inter (chartAt H x).open_source).image_denseRange' - (fun y hy ↦ hasMFDerivWithinAt_extChartAt I hy.2) - fun y hy ↦ ((mdifferentiable_chart _ _).mfderiv_surjective hy.2).denseRange + (fun y hy ↦ hasMFDerivWithinAt_extChartAt hy.2) + fun y hy ↦ ((mdifferentiable_chart _).mfderiv_surjective hy.2).denseRange variable [SmoothManifoldWithCorners I M] in /-- When considering functions between manifolds, this statement shows up often. It entails @@ -107,7 +107,7 @@ theorem UniqueMDiffOn.uniqueDiffOn_inter_preimage (hs : UniqueMDiffOn I s) (x : intro z hz apply (hs z hz.1).inter' apply (hf z hz.1).preimage_mem_nhdsWithin - exact (isOpen_extChartAt_source I' y).mem_nhds hz.2 + exact (isOpen_extChartAt_source y).mem_nhds hz.2 this.uniqueDiffOn_target_inter _ end @@ -120,7 +120,7 @@ variable {F : Type*} [NormedAddCommGroup F] [NormedSpace 𝕜 F] {Z : M → Type theorem Trivialization.mdifferentiable (e : Trivialization F (π F Z)) [MemTrivializationAtlas e] : e.toPartialHomeomorph.MDifferentiable (I.prod 𝓘(𝕜, F)) (I.prod 𝓘(𝕜, F)) := - ⟨(e.smoothOn I).mdifferentiableOn, (e.smoothOn_symm I).mdifferentiableOn⟩ + ⟨e.smoothOn.mdifferentiableOn, e.smoothOn_symm.mdifferentiableOn⟩ theorem UniqueMDiffWithinAt.smooth_bundle_preimage {p : TotalSpace F Z} (hs : UniqueMDiffWithinAt I s p.proj) : diff --git a/Mathlib/Geometry/Manifold/PartitionOfUnity.lean b/Mathlib/Geometry/Manifold/PartitionOfUnity.lean index d0c29cb8d57a6..db39ef95a11c9 100644 --- a/Mathlib/Geometry/Manifold/PartitionOfUnity.lean +++ b/Mathlib/Geometry/Manifold/PartitionOfUnity.lean @@ -371,7 +371,6 @@ theorem IsSubordinate.support_subset {fs : SmoothBumpCovering ι I M s} {U : M Subset.trans subset_closure (h i) variable (I) in - /-- Let `M` be a smooth manifold with corners modelled on a finite dimensional real vector space. Suppose also that `M` is a Hausdorff `σ`-compact topological space. Let `s` be a closed set in `M` and `U : M → Set M` be a collection of sets such that `U x ∈ 𝓝 x` for every `x ∈ s`. @@ -383,7 +382,7 @@ theorem exists_isSubordinate [T2Space M] [SigmaCompactSpace M] (hs : IsClosed s) haveI : LocallyCompactSpace H := I.locallyCompactSpace haveI : LocallyCompactSpace M := ChartedSpace.locallyCompactSpace H M -- Next we choose a covering by supports of smooth bump functions - have hB := fun x hx => SmoothBumpFunction.nhds_basis_support I (hU x hx) + have hB := fun x hx => SmoothBumpFunction.nhds_basis_support (I := I) (hU x hx) rcases refinement_of_locallyCompact_sigmaCompact_of_nhds_basis_set hs hB with ⟨ι, c, f, hf, hsub', hfin⟩ choose hcs hfU using hf @@ -718,7 +717,7 @@ theorem IsOpen.exists_msmooth_support_eq {s : Set M} (hs : IsOpen s) : · apply SmoothPartitionOfUnity.smooth_finsum_smul intro c x hx apply (g_diff c (chartAt H c x)).comp - exact contMDiffAt_of_mem_maximalAtlas (SmoothManifoldWithCorners.chart_mem_maximalAtlas I _) + exact contMDiffAt_of_mem_maximalAtlas (SmoothManifoldWithCorners.chart_mem_maximalAtlas _) (hf c hx) · intro x apply finsum_nonneg (fun c ↦ h''g c x) diff --git a/Mathlib/Geometry/Manifold/Sheaf/Smooth.lean b/Mathlib/Geometry/Manifold/Sheaf/Smooth.lean index d8ac4a298bca8..6cc1c879d0531 100644 --- a/Mathlib/Geometry/Manifold/Sheaf/Smooth.lean +++ b/Mathlib/Geometry/Manifold/Sheaf/Smooth.lean @@ -82,13 +82,13 @@ section TypeCat /-- The sheaf of smooth functions from `M` to `N`, as a sheaf of types. -/ def smoothSheaf : TopCat.Sheaf (Type u) (TopCat.of M) := - (contDiffWithinAt_localInvariantProp IM I ⊤).sheaf M N + (contDiffWithinAt_localInvariantProp (I := IM) (I' := I) ⊤).sheaf M N variable {M} instance smoothSheaf.coeFun (U : (Opens (TopCat.of M))ᵒᵖ) : CoeFun ((smoothSheaf IM I M N).presheaf.obj U) (fun _ ↦ ↑(unop U) → N) := - (contDiffWithinAt_localInvariantProp IM I ⊤).sheafHasCoeToFun _ _ _ + (contDiffWithinAt_localInvariantProp ⊤).sheafHasCoeToFun _ _ _ open Manifold in /-- The object of `smoothSheaf IM I M N` for the open set `U` in `M` is @@ -135,12 +135,12 @@ variable {IM I N} @[simp] lemma smoothSheaf.eval_germ (U : Opens M) (x : M) (hx : x ∈ U) (f : (smoothSheaf IM I M N).presheaf.obj (op U)) : smoothSheaf.eval IM I N (x : M) ((smoothSheaf IM I M N).presheaf.germ U x hx f) = f ⟨x, hx⟩ := - TopCat.stalkToFiber_germ ((contDiffWithinAt_localInvariantProp IM I ⊤).localPredicate M N) _ _ _ _ + TopCat.stalkToFiber_germ ((contDiffWithinAt_localInvariantProp ⊤).localPredicate M N) _ _ _ _ lemma smoothSheaf.smooth_section {U : (Opens (TopCat.of M))ᵒᵖ} (f : (smoothSheaf IM I M N).presheaf.obj U) : Smooth IM I f := - (contDiffWithinAt_localInvariantProp IM I ⊤).section_spec _ _ _ _ + (contDiffWithinAt_localInvariantProp ⊤).section_spec _ _ _ _ end TypeCat @@ -278,7 +278,7 @@ example : (CategoryTheory.sheafCompose _ (CategoryTheory.forget CommRingCat.{u}) instance smoothSheafCommRing.coeFun (U : (Opens (TopCat.of M))ᵒᵖ) : CoeFun ((smoothSheafCommRing IM I M R).presheaf.obj U) (fun _ ↦ ↑(unop U) → R) := - (contDiffWithinAt_localInvariantProp IM I ⊤).sheafHasCoeToFun _ _ _ + (contDiffWithinAt_localInvariantProp ⊤).sheafHasCoeToFun _ _ _ open CategoryTheory Limits diff --git a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean index 85e1eada2814c..f93d14d8c5e1b 100644 --- a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean +++ b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean @@ -496,11 +496,10 @@ section contDiffGroupoid variable {m n : ℕ∞} {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedAddCommGroup E] - [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] (I : ModelWithCorners 𝕜 E H) {M : Type*} + [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] {I : ModelWithCorners 𝕜 E H} {M : Type*} [TopologicalSpace M] -variable (n) - +variable (n I) in /-- Given a model with corners `(E, H)`, we define the pregroupoid of `C^n` transformations of `H` as the maps that are `C^n` when read in `E` through `I`. -/ def contDiffPregroupoid : Pregroupoid H where @@ -539,13 +538,12 @@ def contDiffPregroupoid : Pregroupoid H where simp only [mfld_simps] at hy1 ⊢ rw [fg _ hy1] +variable (n I) in /-- Given a model with corners `(E, H)`, we define the groupoid of invertible `C^n` transformations of `H` as the invertible maps that are `C^n` when read in `E` through `I`. -/ def contDiffGroupoid : StructureGroupoid H := Pregroupoid.groupoid (contDiffPregroupoid n I) -variable {n} - /-- Inclusion of the groupoid of `C^n` local diffeos in the groupoid of `C^m` local diffeos when `m ≤ n` -/ theorem contDiffGroupoid_le (h : m ≤ n) : contDiffGroupoid n I ≤ contDiffGroupoid m I := by @@ -570,8 +568,6 @@ theorem contDiffGroupoid_zero_eq : contDiffGroupoid 0 I = continuousGroupoid H : · refine I.continuous.comp_continuousOn (u.symm.continuousOn.comp I.continuousOn_symm ?_) exact (mapsTo_preimage _ _).mono_left inter_subset_left -variable (n) - /-- An identity partial homeomorphism belongs to the `C^n` groupoid. -/ theorem ofSet_mem_contDiffGroupoid {s : Set H} (hs : IsOpen s) : PartialHomeomorph.ofSet s hs ∈ contDiffGroupoid n I := by @@ -587,7 +583,7 @@ theorem symm_trans_mem_contDiffGroupoid (e : PartialHomeomorph M H) : e.symm.trans e ∈ contDiffGroupoid n I := haveI : e.symm.trans e ≈ PartialHomeomorph.ofSet e.target e.open_target := PartialHomeomorph.symm_trans_self _ - StructureGroupoid.mem_of_eqOnSource _ (ofSet_mem_contDiffGroupoid n I e.open_target) this + StructureGroupoid.mem_of_eqOnSource _ (ofSet_mem_contDiffGroupoid e.open_target) this variable {E' H' : Type*} [NormedAddCommGroup E'] [NormedSpace 𝕜 E'] [TopologicalSpace H'] @@ -616,7 +612,7 @@ instance : ClosedUnderRestriction (contDiffGroupoid n I) := rw [StructureGroupoid.le_iff] rintro e ⟨s, hs, hes⟩ apply (contDiffGroupoid n I).mem_of_eqOnSource' _ _ _ hes - exact ofSet_mem_contDiffGroupoid n I hs) + exact ofSet_mem_contDiffGroupoid hs) end contDiffGroupoid @@ -670,7 +666,7 @@ model with corners `I`. -/ def maximalAtlas := (contDiffGroupoid ∞ I).maximalAtlas M -variable {M} +variable {I M} theorem subset_maximalAtlas [SmoothManifoldWithCorners I M] : atlas H M ⊆ maximalAtlas I M := StructureGroupoid.subset_maximalAtlas _ @@ -679,8 +675,6 @@ theorem chart_mem_maximalAtlas [SmoothManifoldWithCorners I M] (x : M) : chartAt H x ∈ maximalAtlas I M := StructureGroupoid.chart_mem_maximalAtlas _ x -variable {I} - theorem compatible_of_mem_maximalAtlas {e e' : PartialHomeomorph M H} (he : e ∈ maximalAtlas I M) (he' : e' ∈ maximalAtlas I M) : e.symm.trans e' ∈ contDiffGroupoid ∞ I := StructureGroupoid.compatible_of_mem_maximalAtlas he he' @@ -721,7 +715,7 @@ end SmoothManifoldWithCorners theorem PartialHomeomorph.singleton_smoothManifoldWithCorners {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] - {H : Type*} [TopologicalSpace H] (I : ModelWithCorners 𝕜 E H) + {H : Type*} [TopologicalSpace H] {I : ModelWithCorners 𝕜 E H} {M : Type*} [TopologicalSpace M] (e : PartialHomeomorph M H) (h : e.source = Set.univ) : @SmoothManifoldWithCorners 𝕜 _ E _ _ H _ I M _ (e.singletonChartedSpace h) := @SmoothManifoldWithCorners.mk' _ _ _ _ _ _ _ _ _ _ (id _) <| @@ -729,10 +723,10 @@ theorem PartialHomeomorph.singleton_smoothManifoldWithCorners theorem IsOpenEmbedding.singleton_smoothManifoldWithCorners {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] - (I : ModelWithCorners 𝕜 E H) {M : Type*} [TopologicalSpace M] [Nonempty M] {f : M → H} + {I : ModelWithCorners 𝕜 E H} {M : Type*} [TopologicalSpace M] [Nonempty M] {f : M → H} (h : IsOpenEmbedding f) : @SmoothManifoldWithCorners 𝕜 _ E _ _ H _ I M _ h.singletonChartedSpace := - (h.toPartialHomeomorph f).singleton_smoothManifoldWithCorners I (by simp) + (h.toPartialHomeomorph f).singleton_smoothManifoldWithCorners (by simp) @[deprecated (since := "2024-10-18")] alias OpenEmbedding.singleton_smoothManifoldWithCorners := @@ -743,7 +737,7 @@ namespace TopologicalSpace.Opens open TopologicalSpace variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedAddCommGroup E] - [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] (I : ModelWithCorners 𝕜 E H) {M : Type*} + [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] {I : ModelWithCorners 𝕜 E H} {M : Type*} [TopologicalSpace M] [ChartedSpace H M] [SmoothManifoldWithCorners I M] (s : Opens M) instance : SmoothManifoldWithCorners I s := @@ -757,8 +751,8 @@ open scoped Topology variable {𝕜 E M H E' M' H' : Type*} [NontriviallyNormedField 𝕜] [NormedAddCommGroup E] [NormedSpace 𝕜 E] [TopologicalSpace H] [TopologicalSpace M] (f f' : PartialHomeomorph M H) - (I : ModelWithCorners 𝕜 E H) [NormedAddCommGroup E'] [NormedSpace 𝕜 E'] [TopologicalSpace H'] - [TopologicalSpace M'] (I' : ModelWithCorners 𝕜 E' H') {s t : Set M} + {I : ModelWithCorners 𝕜 E H} [NormedAddCommGroup E'] [NormedSpace 𝕜 E'] [TopologicalSpace H'] + [TopologicalSpace M'] {I' : ModelWithCorners 𝕜 E' H'} {s t : Set M} /-! ### Extended charts @@ -772,6 +766,7 @@ as `PartialEquiv`. namespace PartialHomeomorph +variable (I) in /-- Given a chart `f` on a manifold with corners, `f.extend I` is the extended chart to the model vector space. -/ @[simp, mfld_simps] @@ -812,13 +807,13 @@ theorem extend_left_inv {x : M} (hxf : x ∈ f.source) : (f.extend I).symm (f.ex /-- Variant of `f.extend_left_inv I`, stated in terms of images. -/ lemma extend_left_inv' (ht : t ⊆ f.source) : ((f.extend I).symm ∘ (f.extend I)) '' t = t := - EqOn.image_eq_self (fun _ hx ↦ f.extend_left_inv I (ht hx)) + EqOn.image_eq_self (fun _ hx ↦ f.extend_left_inv (ht hx)) theorem extend_source_mem_nhds {x : M} (h : x ∈ f.source) : (f.extend I).source ∈ 𝓝 x := - (isOpen_extend_source f I).mem_nhds <| by rwa [f.extend_source I] + (isOpen_extend_source f).mem_nhds <| by rwa [f.extend_source] theorem extend_source_mem_nhdsWithin {x : M} (h : x ∈ f.source) : (f.extend I).source ∈ 𝓝[s] x := - mem_nhdsWithin_of_mem_nhds <| extend_source_mem_nhds f I h + mem_nhdsWithin_of_mem_nhds <| extend_source_mem_nhds f h theorem continuousOn_extend : ContinuousOn (f.extend I) (f.extend I).source := by refine I.continuous.comp_continuousOn ?_ @@ -826,7 +821,7 @@ theorem continuousOn_extend : ContinuousOn (f.extend I) (f.extend I).source := b exact f.continuousOn theorem continuousAt_extend {x : M} (h : x ∈ f.source) : ContinuousAt (f.extend I) x := - (continuousOn_extend f I).continuousAt <| extend_source_mem_nhds f I h + (continuousOn_extend f).continuousAt <| extend_source_mem_nhds f h theorem map_extend_nhds {x : M} (hy : x ∈ f.source) : map (f.extend I) (𝓝 x) = 𝓝[range I] f.extend I x := by @@ -834,16 +829,16 @@ theorem map_extend_nhds {x : M} (hy : x ∈ f.source) : theorem map_extend_nhds_of_boundaryless [I.Boundaryless] {x : M} (hx : x ∈ f.source) : map (f.extend I) (𝓝 x) = 𝓝 (f.extend I x) := by - rw [f.map_extend_nhds _ hx, I.range_eq_univ, nhdsWithin_univ] + rw [f.map_extend_nhds hx, I.range_eq_univ, nhdsWithin_univ] theorem extend_target_mem_nhdsWithin {y : M} (hy : y ∈ f.source) : (f.extend I).target ∈ 𝓝[range I] f.extend I y := by - rw [← PartialEquiv.image_source_eq_target, ← map_extend_nhds f I hy] - exact image_mem_map (extend_source_mem_nhds _ _ hy) + rw [← PartialEquiv.image_source_eq_target, ← map_extend_nhds f hy] + exact image_mem_map (extend_source_mem_nhds _ hy) theorem extend_image_nhd_mem_nhds_of_boundaryless [I.Boundaryless] {x} (hx : x ∈ f.source) {s : Set M} (h : s ∈ 𝓝 x) : (f.extend I) '' s ∈ 𝓝 ((f.extend I) x) := by - rw [← f.map_extend_nhds_of_boundaryless _ hx, Filter.mem_map] + rw [← f.map_extend_nhds_of_boundaryless hx, Filter.mem_map] filter_upwards [h] using subset_preimage_image (f.extend I) s theorem extend_target_subset_range : (f.extend I).target ⊆ range I := by simp only [mfld_simps] @@ -863,8 +858,8 @@ lemma mem_interior_extend_target {y : H} (hy : y ∈ f.target) theorem nhdsWithin_extend_target_eq {y : M} (hy : y ∈ f.source) : 𝓝[(f.extend I).target] f.extend I y = 𝓝[range I] f.extend I y := - (nhdsWithin_mono _ (extend_target_subset_range _ _)).antisymm <| - nhdsWithin_le_of_mem (extend_target_mem_nhdsWithin _ _ hy) + (nhdsWithin_mono _ (extend_target_subset_range _)).antisymm <| + nhdsWithin_le_of_mem (extend_target_mem_nhdsWithin _ hy) theorem continuousAt_extend_symm' {x : E} (h : x ∈ (f.extend I).target) : ContinuousAt (f.extend I).symm x := @@ -872,10 +867,10 @@ theorem continuousAt_extend_symm' {x : E} (h : x ∈ (f.extend I).target) : theorem continuousAt_extend_symm {x : M} (h : x ∈ f.source) : ContinuousAt (f.extend I).symm (f.extend I x) := - continuousAt_extend_symm' f I <| (f.extend I).map_source <| by rwa [f.extend_source] + continuousAt_extend_symm' f <| (f.extend I).map_source <| by rwa [f.extend_source] theorem continuousOn_extend_symm : ContinuousOn (f.extend I).symm (f.extend I).target := fun _ h => - (continuousAt_extend_symm' _ _ h).continuousWithinAt + (continuousAt_extend_symm' _ h).continuousWithinAt theorem extend_symm_continuousWithinAt_comp_right_iff {X} [TopologicalSpace X] {g : M → X} {s : Set M} {x : M} : @@ -885,52 +880,52 @@ theorem extend_symm_continuousWithinAt_comp_right_iff {X} [TopologicalSpace X] { theorem isOpen_extend_preimage' {s : Set E} (hs : IsOpen s) : IsOpen ((f.extend I).source ∩ f.extend I ⁻¹' s) := - (continuousOn_extend f I).isOpen_inter_preimage (isOpen_extend_source _ _) hs + (continuousOn_extend f).isOpen_inter_preimage (isOpen_extend_source _) hs theorem isOpen_extend_preimage {s : Set E} (hs : IsOpen s) : IsOpen (f.source ∩ f.extend I ⁻¹' s) := by - rw [← extend_source f I]; exact isOpen_extend_preimage' f I hs + rw [← extend_source f (I := I)]; exact isOpen_extend_preimage' f hs theorem map_extend_nhdsWithin_eq_image {y : M} (hy : y ∈ f.source) : map (f.extend I) (𝓝[s] y) = 𝓝[f.extend I '' ((f.extend I).source ∩ s)] f.extend I y := by set e := f.extend I calc map e (𝓝[s] y) = map e (𝓝[e.source ∩ s] y) := - congr_arg (map e) (nhdsWithin_inter_of_mem (extend_source_mem_nhdsWithin f I hy)).symm + congr_arg (map e) (nhdsWithin_inter_of_mem (extend_source_mem_nhdsWithin f hy)).symm _ = 𝓝[e '' (e.source ∩ s)] e y := ((f.extend I).leftInvOn.mono inter_subset_left).map_nhdsWithin_eq ((f.extend I).left_inv <| by rwa [f.extend_source]) - (continuousAt_extend_symm f I hy).continuousWithinAt - (continuousAt_extend f I hy).continuousWithinAt + (continuousAt_extend_symm f hy).continuousWithinAt + (continuousAt_extend f hy).continuousWithinAt theorem map_extend_nhdsWithin_eq_image_of_subset {y : M} (hy : y ∈ f.source) (hs : s ⊆ f.source) : map (f.extend I) (𝓝[s] y) = 𝓝[f.extend I '' s] f.extend I y := by - rw [map_extend_nhdsWithin_eq_image _ _ hy, inter_eq_self_of_subset_right] + rw [map_extend_nhdsWithin_eq_image _ hy, inter_eq_self_of_subset_right] rwa [extend_source] theorem map_extend_nhdsWithin {y : M} (hy : y ∈ f.source) : map (f.extend I) (𝓝[s] y) = 𝓝[(f.extend I).symm ⁻¹' s ∩ range I] f.extend I y := by - rw [map_extend_nhdsWithin_eq_image f I hy, nhdsWithin_inter, ← - nhdsWithin_extend_target_eq _ _ hy, ← nhdsWithin_inter, (f.extend I).image_source_inter_eq', + rw [map_extend_nhdsWithin_eq_image f hy, nhdsWithin_inter, ← + nhdsWithin_extend_target_eq _ hy, ← nhdsWithin_inter, (f.extend I).image_source_inter_eq', inter_comm] theorem map_extend_symm_nhdsWithin {y : M} (hy : y ∈ f.source) : map (f.extend I).symm (𝓝[(f.extend I).symm ⁻¹' s ∩ range I] f.extend I y) = 𝓝[s] y := by - rw [← map_extend_nhdsWithin f I hy, map_map, Filter.map_congr, map_id] - exact (f.extend I).leftInvOn.eqOn.eventuallyEq_of_mem (extend_source_mem_nhdsWithin _ _ hy) + rw [← map_extend_nhdsWithin f hy, map_map, Filter.map_congr, map_id] + exact (f.extend I).leftInvOn.eqOn.eventuallyEq_of_mem (extend_source_mem_nhdsWithin _ hy) theorem map_extend_symm_nhdsWithin_range {y : M} (hy : y ∈ f.source) : map (f.extend I).symm (𝓝[range I] f.extend I y) = 𝓝 y := by - rw [← nhdsWithin_univ, ← map_extend_symm_nhdsWithin f I hy, preimage_univ, univ_inter] + rw [← nhdsWithin_univ, ← map_extend_symm_nhdsWithin f (I := I) hy, preimage_univ, univ_inter] theorem tendsto_extend_comp_iff {α : Type*} {l : Filter α} {g : α → M} (hg : ∀ᶠ z in l, g z ∈ f.source) {y : M} (hy : y ∈ f.source) : Tendsto (f.extend I ∘ g) l (𝓝 (f.extend I y)) ↔ Tendsto g l (𝓝 y) := by - refine ⟨fun h u hu ↦ mem_map.2 ?_, (continuousAt_extend _ _ hy).tendsto.comp⟩ - have := (f.continuousAt_extend_symm I hy).tendsto.comp h - rw [extend_left_inv _ _ hy] at this + refine ⟨fun h u hu ↦ mem_map.2 ?_, (continuousAt_extend _ hy).tendsto.comp⟩ + have := (f.continuousAt_extend_symm hy).tendsto.comp h + rw [extend_left_inv _ hy] at this filter_upwards [hg, mem_map.1 (this hu)] with z hz hzu - simpa only [(· ∘ ·), extend_left_inv _ _ hz, mem_preimage] using hzu + simpa only [(· ∘ ·), extend_left_inv _ hz, mem_preimage] using hzu -- there is no definition `writtenInExtend` but we already use some made-up names in this file theorem continuousWithinAt_writtenInExtend_iff {f' : PartialHomeomorph M' H'} {g : M → M'} {y : M} @@ -939,11 +934,11 @@ theorem continuousWithinAt_writtenInExtend_iff {f' : PartialHomeomorph M' H'} {g ((f.extend I).symm ⁻¹' s ∩ range I) (f.extend I y) ↔ ContinuousWithinAt g s y := by unfold ContinuousWithinAt simp only [comp_apply] - rw [extend_left_inv _ _ hy, f'.tendsto_extend_comp_iff _ _ hgy, - ← f.map_extend_symm_nhdsWithin I hy, tendsto_map'_iff] - rw [← f.map_extend_nhdsWithin I hy, eventually_map] + rw [extend_left_inv _ hy, f'.tendsto_extend_comp_iff _ hgy, + ← f.map_extend_symm_nhdsWithin (I := I) hy, tendsto_map'_iff] + rw [← f.map_extend_nhdsWithin (I := I) hy, eventually_map] filter_upwards [inter_mem_nhdsWithin _ (f.open_source.mem_nhds hy)] with z hz - rw [comp_apply, extend_left_inv _ _ hz.2] + rw [comp_apply, extend_left_inv _ hz.2] exact hmaps hz.1 -- there is no definition `writtenInExtend` but we already use some made-up names in this file @@ -955,7 +950,7 @@ theorem continuousOn_writtenInExtend_iff {f' : PartialHomeomorph M' H'} {g : M ContinuousOn (f'.extend I' ∘ g ∘ (f.extend I).symm) (f.extend I '' s) ↔ ContinuousOn g s := by refine forall_mem_image.trans <| forall₂_congr fun x hx ↦ ?_ refine (continuousWithinAt_congr_nhds ?_).trans - (continuousWithinAt_writtenInExtend_iff _ _ _ (hs hx) (hmaps hx) hmaps) + (continuousWithinAt_writtenInExtend_iff _ (hs hx) (hmaps hx) hmaps) rw [← map_extend_nhdsWithin_eq_image_of_subset, ← map_extend_nhdsWithin] exacts [hs hx, hs hx, hs] @@ -963,11 +958,11 @@ theorem continuousOn_writtenInExtend_iff {f' : PartialHomeomorph M' H'} {g : M in the source is a neighborhood of the preimage, within a set. -/ theorem extend_preimage_mem_nhdsWithin {x : M} (h : x ∈ f.source) (ht : t ∈ 𝓝[s] x) : (f.extend I).symm ⁻¹' t ∈ 𝓝[(f.extend I).symm ⁻¹' s ∩ range I] f.extend I x := by - rwa [← map_extend_symm_nhdsWithin f I h, mem_map] at ht + rwa [← map_extend_symm_nhdsWithin f (I := I) h, mem_map] at ht theorem extend_preimage_mem_nhds {x : M} (h : x ∈ f.source) (ht : t ∈ 𝓝 x) : (f.extend I).symm ⁻¹' t ∈ 𝓝 (f.extend I x) := by - apply (continuousAt_extend_symm f I h).preimage_mem_nhds + apply (continuousAt_extend_symm f h).preimage_mem_nhds rwa [(f.extend I).left_inv] rwa [f.extend_source] @@ -994,8 +989,8 @@ theorem extend_symm_preimage_inter_range_eventuallyEq_aux {s : Set M} {x : M} (h theorem extend_symm_preimage_inter_range_eventuallyEq {s : Set M} {x : M} (hs : s ⊆ f.source) (hx : x ∈ f.source) : ((f.extend I).symm ⁻¹' s ∩ range I : Set _) =ᶠ[𝓝 (f.extend I x)] f.extend I '' s := by - rw [← nhdsWithin_eq_iff_eventuallyEq, ← map_extend_nhdsWithin _ _ hx, - map_extend_nhdsWithin_eq_image_of_subset _ _ hx hs] + rw [← nhdsWithin_eq_iff_eventuallyEq, ← map_extend_nhdsWithin _ hx, + map_extend_nhdsWithin_eq_image_of_subset _ hx hs] /-! We use the name `extend_coord_change` for `(f'.extend I).symm ≫ f.extend I`. -/ @@ -1038,7 +1033,7 @@ theorem contDiffOn_extend_coord_change [ChartedSpace H M] (hf : f ∈ maximalAtl theorem contDiffWithinAt_extend_coord_change [ChartedSpace H M] (hf : f ∈ maximalAtlas I M) (hf' : f' ∈ maximalAtlas I M) {x : E} (hx : x ∈ ((f'.extend I).symm ≫ f.extend I).source) : ContDiffWithinAt 𝕜 ⊤ (f.extend I ∘ (f'.extend I).symm) (range I) x := by - apply (contDiffOn_extend_coord_change I hf hf' x hx).mono_of_mem + apply (contDiffOn_extend_coord_change hf hf' x hx).mono_of_mem rw [extend_coord_change_source] at hx ⊢ obtain ⟨z, hz, rfl⟩ := hx exact I.image_mem_nhdsWithin ((PartialHomeomorph.open_source _).mem_nhds hz) @@ -1046,7 +1041,7 @@ theorem contDiffWithinAt_extend_coord_change [ChartedSpace H M] (hf : f ∈ maxi theorem contDiffWithinAt_extend_coord_change' [ChartedSpace H M] (hf : f ∈ maximalAtlas I M) (hf' : f' ∈ maximalAtlas I M) {x : M} (hxf : x ∈ f.source) (hxf' : x ∈ f'.source) : ContDiffWithinAt 𝕜 ⊤ (f.extend I ∘ (f'.extend I).symm) (range I) (f'.extend I x) := by - refine contDiffWithinAt_extend_coord_change I hf hf' ?_ + refine contDiffWithinAt_extend_coord_change hf hf' ?_ rw [← extend_image_source_inter] exact mem_image_of_mem _ ⟨hxf', hxf⟩ @@ -1056,6 +1051,7 @@ open PartialHomeomorph variable [ChartedSpace H M] [ChartedSpace H' M'] +variable (I) in /-- The preferred extended chart on a manifold with corners around a point `x`, from a neighborhood of `x` to the model vector space. -/ @[simp, mfld_simps] @@ -1068,21 +1064,23 @@ theorem extChartAt_coe (x : M) : ⇑(extChartAt I x) = I ∘ chartAt H x := theorem extChartAt_coe_symm (x : M) : ⇑(extChartAt I x).symm = (chartAt H x).symm ∘ I.symm := rfl +variable (I) in theorem extChartAt_source (x : M) : (extChartAt I x).source = (chartAt H x).source := - extend_source _ _ + extend_source _ theorem isOpen_extChartAt_source (x : M) : IsOpen (extChartAt I x).source := - isOpen_extend_source _ _ + isOpen_extend_source _ theorem mem_extChartAt_source (x : M) : x ∈ (extChartAt I x).source := by simp only [extChartAt_source, mem_chart_source] theorem mem_extChartAt_target (x : M) : extChartAt I x x ∈ (extChartAt I x).target := - (extChartAt I x).map_source <| mem_extChartAt_source _ _ + (extChartAt I x).map_source <| mem_extChartAt_source _ +variable (I) in theorem extChartAt_target (x : M) : (extChartAt I x).target = I.symm ⁻¹' (chartAt H x).target ∩ range I := - extend_target _ _ + extend_target _ theorem uniqueDiffOn_extChartAt_target (x : M) : UniqueDiffOn 𝕜 (extChartAt I x).target := by rw [extChartAt_target] @@ -1090,64 +1088,64 @@ theorem uniqueDiffOn_extChartAt_target (x : M) : UniqueDiffOn 𝕜 (extChartAt I theorem uniqueDiffWithinAt_extChartAt_target (x : M) : UniqueDiffWithinAt 𝕜 (extChartAt I x).target (extChartAt I x x) := - uniqueDiffOn_extChartAt_target I x _ <| mem_extChartAt_target I x + uniqueDiffOn_extChartAt_target x _ <| mem_extChartAt_target x theorem extChartAt_to_inv (x : M) : (extChartAt I x).symm ((extChartAt I x) x) = x := - (extChartAt I x).left_inv (mem_extChartAt_source I x) + (extChartAt I x).left_inv (mem_extChartAt_source x) theorem mapsTo_extChartAt {x : M} (hs : s ⊆ (chartAt H x).source) : MapsTo (extChartAt I x) s ((extChartAt I x).symm ⁻¹' s ∩ range I) := - mapsTo_extend _ _ hs + mapsTo_extend _ hs theorem extChartAt_source_mem_nhds' {x x' : M} (h : x' ∈ (extChartAt I x).source) : (extChartAt I x).source ∈ 𝓝 x' := - extend_source_mem_nhds _ _ <| by rwa [← extChartAt_source I] + extend_source_mem_nhds _ <| by rwa [← extChartAt_source I] theorem extChartAt_source_mem_nhds (x : M) : (extChartAt I x).source ∈ 𝓝 x := - extChartAt_source_mem_nhds' I (mem_extChartAt_source I x) + extChartAt_source_mem_nhds' (mem_extChartAt_source x) theorem extChartAt_source_mem_nhdsWithin' {x x' : M} (h : x' ∈ (extChartAt I x).source) : (extChartAt I x).source ∈ 𝓝[s] x' := - mem_nhdsWithin_of_mem_nhds (extChartAt_source_mem_nhds' I h) + mem_nhdsWithin_of_mem_nhds (extChartAt_source_mem_nhds' h) theorem extChartAt_source_mem_nhdsWithin (x : M) : (extChartAt I x).source ∈ 𝓝[s] x := - mem_nhdsWithin_of_mem_nhds (extChartAt_source_mem_nhds I x) + mem_nhdsWithin_of_mem_nhds (extChartAt_source_mem_nhds x) theorem continuousOn_extChartAt (x : M) : ContinuousOn (extChartAt I x) (extChartAt I x).source := - continuousOn_extend _ _ + continuousOn_extend _ theorem continuousAt_extChartAt' {x x' : M} (h : x' ∈ (extChartAt I x).source) : ContinuousAt (extChartAt I x) x' := - continuousAt_extend _ _ <| by rwa [← extChartAt_source I] + continuousAt_extend _ <| by rwa [← extChartAt_source I] theorem continuousAt_extChartAt (x : M) : ContinuousAt (extChartAt I x) x := - continuousAt_extChartAt' _ (mem_extChartAt_source I x) + continuousAt_extChartAt' (mem_extChartAt_source x) theorem map_extChartAt_nhds' {x y : M} (hy : y ∈ (extChartAt I x).source) : map (extChartAt I x) (𝓝 y) = 𝓝[range I] extChartAt I x y := - map_extend_nhds _ _ <| by rwa [← extChartAt_source I] + map_extend_nhds _ <| by rwa [← extChartAt_source I] theorem map_extChartAt_nhds (x : M) : map (extChartAt I x) (𝓝 x) = 𝓝[range I] extChartAt I x x := - map_extChartAt_nhds' I <| mem_extChartAt_source I x + map_extChartAt_nhds' <| mem_extChartAt_source x theorem map_extChartAt_nhds_of_boundaryless [I.Boundaryless] (x : M) : map (extChartAt I x) (𝓝 x) = 𝓝 (extChartAt I x x) := by rw [extChartAt] - exact map_extend_nhds_of_boundaryless (chartAt H x) I (mem_chart_source H x) + exact map_extend_nhds_of_boundaryless (chartAt H x) (mem_chart_source H x) variable {x} in theorem extChartAt_image_nhd_mem_nhds_of_boundaryless [I.Boundaryless] {x : M} (hx : s ∈ 𝓝 x) : extChartAt I x '' s ∈ 𝓝 (extChartAt I x x) := by rw [extChartAt] - exact extend_image_nhd_mem_nhds_of_boundaryless _ I (mem_chart_source H x) hx + exact extend_image_nhd_mem_nhds_of_boundaryless _ (mem_chart_source H x) hx theorem extChartAt_target_mem_nhdsWithin' {x y : M} (hy : y ∈ (extChartAt I x).source) : (extChartAt I x).target ∈ 𝓝[range I] extChartAt I x y := - extend_target_mem_nhdsWithin _ _ <| by rwa [← extChartAt_source I] + extend_target_mem_nhdsWithin _ <| by rwa [← extChartAt_source I] theorem extChartAt_target_mem_nhdsWithin (x : M) : (extChartAt I x).target ∈ 𝓝[range I] extChartAt I x x := - extChartAt_target_mem_nhdsWithin' I (mem_extChartAt_source I x) + extChartAt_target_mem_nhdsWithin' (mem_extChartAt_source x) /-- If we're boundaryless, `extChartAt` has open target -/ theorem isOpen_extChartAt_target [I.Boundaryless] (x : M) : IsOpen (extChartAt I x).target := by @@ -1157,109 +1155,109 @@ theorem isOpen_extChartAt_target [I.Boundaryless] (x : M) : IsOpen (extChartAt I /-- If we're boundaryless, `(extChartAt I x).target` is a neighborhood of the key point -/ theorem extChartAt_target_mem_nhds [I.Boundaryless] (x : M) : (extChartAt I x).target ∈ 𝓝 (extChartAt I x x) := by - convert extChartAt_target_mem_nhdsWithin I x + convert extChartAt_target_mem_nhdsWithin x simp only [I.range_eq_univ, nhdsWithin_univ] /-- If we're boundaryless, `(extChartAt I x).target` is a neighborhood of any of its points -/ theorem extChartAt_target_mem_nhds' [I.Boundaryless] {x : M} {y : E} (m : y ∈ (extChartAt I x).target) : (extChartAt I x).target ∈ 𝓝 y := - (isOpen_extChartAt_target I x).mem_nhds m + (isOpen_extChartAt_target x).mem_nhds m theorem extChartAt_target_subset_range (x : M) : (extChartAt I x).target ⊆ range I := by simp only [mfld_simps] theorem nhdsWithin_extChartAt_target_eq' {x y : M} (hy : y ∈ (extChartAt I x).source) : 𝓝[(extChartAt I x).target] extChartAt I x y = 𝓝[range I] extChartAt I x y := - nhdsWithin_extend_target_eq _ _ <| by rwa [← extChartAt_source I] + nhdsWithin_extend_target_eq _ <| by rwa [← extChartAt_source I] theorem nhdsWithin_extChartAt_target_eq (x : M) : 𝓝[(extChartAt I x).target] (extChartAt I x) x = 𝓝[range I] (extChartAt I x) x := - nhdsWithin_extChartAt_target_eq' I (mem_extChartAt_source I x) + nhdsWithin_extChartAt_target_eq' (mem_extChartAt_source x) theorem continuousAt_extChartAt_symm'' {x : M} {y : E} (h : y ∈ (extChartAt I x).target) : ContinuousAt (extChartAt I x).symm y := - continuousAt_extend_symm' _ _ h + continuousAt_extend_symm' _ h theorem continuousAt_extChartAt_symm' {x x' : M} (h : x' ∈ (extChartAt I x).source) : ContinuousAt (extChartAt I x).symm (extChartAt I x x') := - continuousAt_extChartAt_symm'' I <| (extChartAt I x).map_source h + continuousAt_extChartAt_symm'' <| (extChartAt I x).map_source h theorem continuousAt_extChartAt_symm (x : M) : ContinuousAt (extChartAt I x).symm ((extChartAt I x) x) := - continuousAt_extChartAt_symm' I (mem_extChartAt_source I x) + continuousAt_extChartAt_symm' (mem_extChartAt_source x) theorem continuousOn_extChartAt_symm (x : M) : ContinuousOn (extChartAt I x).symm (extChartAt I x).target := - fun _y hy => (continuousAt_extChartAt_symm'' _ hy).continuousWithinAt + fun _y hy => (continuousAt_extChartAt_symm'' hy).continuousWithinAt theorem isOpen_extChartAt_preimage' (x : M) {s : Set E} (hs : IsOpen s) : IsOpen ((extChartAt I x).source ∩ extChartAt I x ⁻¹' s) := - isOpen_extend_preimage' _ _ hs + isOpen_extend_preimage' _ hs theorem isOpen_extChartAt_preimage (x : M) {s : Set E} (hs : IsOpen s) : IsOpen ((chartAt H x).source ∩ extChartAt I x ⁻¹' s) := by rw [← extChartAt_source I] - exact isOpen_extChartAt_preimage' I x hs + exact isOpen_extChartAt_preimage' x hs theorem map_extChartAt_nhdsWithin_eq_image' {x y : M} (hy : y ∈ (extChartAt I x).source) : map (extChartAt I x) (𝓝[s] y) = 𝓝[extChartAt I x '' ((extChartAt I x).source ∩ s)] extChartAt I x y := - map_extend_nhdsWithin_eq_image _ _ <| by rwa [← extChartAt_source I] + map_extend_nhdsWithin_eq_image _ <| by rwa [← extChartAt_source I] theorem map_extChartAt_nhdsWithin_eq_image (x : M) : map (extChartAt I x) (𝓝[s] x) = 𝓝[extChartAt I x '' ((extChartAt I x).source ∩ s)] extChartAt I x x := - map_extChartAt_nhdsWithin_eq_image' I (mem_extChartAt_source I x) + map_extChartAt_nhdsWithin_eq_image' (mem_extChartAt_source x) theorem map_extChartAt_nhdsWithin' {x y : M} (hy : y ∈ (extChartAt I x).source) : map (extChartAt I x) (𝓝[s] y) = 𝓝[(extChartAt I x).symm ⁻¹' s ∩ range I] extChartAt I x y := - map_extend_nhdsWithin _ _ <| by rwa [← extChartAt_source I] + map_extend_nhdsWithin _ <| by rwa [← extChartAt_source I] theorem map_extChartAt_nhdsWithin (x : M) : map (extChartAt I x) (𝓝[s] x) = 𝓝[(extChartAt I x).symm ⁻¹' s ∩ range I] extChartAt I x x := - map_extChartAt_nhdsWithin' I (mem_extChartAt_source I x) + map_extChartAt_nhdsWithin' (mem_extChartAt_source x) theorem map_extChartAt_symm_nhdsWithin' {x y : M} (hy : y ∈ (extChartAt I x).source) : map (extChartAt I x).symm (𝓝[(extChartAt I x).symm ⁻¹' s ∩ range I] extChartAt I x y) = 𝓝[s] y := - map_extend_symm_nhdsWithin _ _ <| by rwa [← extChartAt_source I] + map_extend_symm_nhdsWithin _ <| by rwa [← extChartAt_source I] theorem map_extChartAt_symm_nhdsWithin_range' {x y : M} (hy : y ∈ (extChartAt I x).source) : map (extChartAt I x).symm (𝓝[range I] extChartAt I x y) = 𝓝 y := - map_extend_symm_nhdsWithin_range _ _ <| by rwa [← extChartAt_source I] + map_extend_symm_nhdsWithin_range _ <| by rwa [← extChartAt_source I] theorem map_extChartAt_symm_nhdsWithin (x : M) : map (extChartAt I x).symm (𝓝[(extChartAt I x).symm ⁻¹' s ∩ range I] extChartAt I x x) = 𝓝[s] x := - map_extChartAt_symm_nhdsWithin' I (mem_extChartAt_source I x) + map_extChartAt_symm_nhdsWithin' (mem_extChartAt_source x) theorem map_extChartAt_symm_nhdsWithin_range (x : M) : map (extChartAt I x).symm (𝓝[range I] extChartAt I x x) = 𝓝 x := - map_extChartAt_symm_nhdsWithin_range' I (mem_extChartAt_source I x) + map_extChartAt_symm_nhdsWithin_range' (mem_extChartAt_source x) /-- Technical lemma ensuring that the preimage under an extended chart of a neighborhood of a point in the source is a neighborhood of the preimage, within a set. -/ theorem extChartAt_preimage_mem_nhdsWithin' {x x' : M} (h : x' ∈ (extChartAt I x).source) (ht : t ∈ 𝓝[s] x') : (extChartAt I x).symm ⁻¹' t ∈ 𝓝[(extChartAt I x).symm ⁻¹' s ∩ range I] (extChartAt I x) x' := by - rwa [← map_extChartAt_symm_nhdsWithin' I h, mem_map] at ht + rwa [← map_extChartAt_symm_nhdsWithin' h, mem_map] at ht /-- Technical lemma ensuring that the preimage under an extended chart of a neighborhood of the base point is a neighborhood of the preimage, within a set. -/ theorem extChartAt_preimage_mem_nhdsWithin {x : M} (ht : t ∈ 𝓝[s] x) : (extChartAt I x).symm ⁻¹' t ∈ 𝓝[(extChartAt I x).symm ⁻¹' s ∩ range I] (extChartAt I x) x := - extChartAt_preimage_mem_nhdsWithin' I (mem_extChartAt_source I x) ht + extChartAt_preimage_mem_nhdsWithin' (mem_extChartAt_source x) ht theorem extChartAt_preimage_mem_nhds' {x x' : M} (h : x' ∈ (extChartAt I x).source) (ht : t ∈ 𝓝 x') : (extChartAt I x).symm ⁻¹' t ∈ 𝓝 (extChartAt I x x') := - extend_preimage_mem_nhds _ _ (by rwa [← extChartAt_source I]) ht + extend_preimage_mem_nhds _ (by rwa [← extChartAt_source I]) ht /-- Technical lemma ensuring that the preimage under an extended chart of a neighborhood of a point is a neighborhood of the preimage. -/ theorem extChartAt_preimage_mem_nhds {x : M} (ht : t ∈ 𝓝 x) : (extChartAt I x).symm ⁻¹' t ∈ 𝓝 ((extChartAt I x) x) := by - apply (continuousAt_extChartAt_symm I x).preimage_mem_nhds - rwa [(extChartAt I x).left_inv (mem_extChartAt_source _ _)] + apply (continuousAt_extChartAt_symm x).preimage_mem_nhds + rwa [(extChartAt I x).left_inv (mem_extChartAt_source _)] /-- Technical lemma to rewrite suitably the preimage of an intersection under an extended chart, to bring it into a convenient form to apply derivative lemmas. -/ @@ -1275,28 +1273,29 @@ theorem ContinuousWithinAt.nhdsWithin_extChartAt_symm_preimage_inter_range (extChartAt I x).symm ⁻¹' (s ∩ f ⁻¹' (extChartAt I' (f x)).source)] (extChartAt I x x) := by rw [← (extChartAt I x).image_source_inter_eq', ← map_extChartAt_nhdsWithin_eq_image, ← map_extChartAt_nhdsWithin, nhdsWithin_inter_of_mem'] - exact hc (extChartAt_source_mem_nhds _ _) + exact hc (extChartAt_source_mem_nhds _) /-! We use the name `ext_coord_change` for `(extChartAt I x').symm ≫ extChartAt I x`. -/ theorem ext_coord_change_source (x x' : M) : ((extChartAt I x').symm ≫ extChartAt I x).source = I '' ((chartAt H x').symm ≫ₕ chartAt H x).source := - extend_coord_change_source _ _ _ + extend_coord_change_source _ _ open SmoothManifoldWithCorners theorem contDiffOn_ext_coord_change [SmoothManifoldWithCorners I M] (x x' : M) : ContDiffOn 𝕜 ⊤ (extChartAt I x ∘ (extChartAt I x').symm) ((extChartAt I x').symm ≫ extChartAt I x).source := - contDiffOn_extend_coord_change I (chart_mem_maximalAtlas I x) (chart_mem_maximalAtlas I x') + contDiffOn_extend_coord_change (chart_mem_maximalAtlas x) (chart_mem_maximalAtlas x') theorem contDiffWithinAt_ext_coord_change [SmoothManifoldWithCorners I M] (x x' : M) {y : E} (hy : y ∈ ((extChartAt I x').symm ≫ extChartAt I x).source) : ContDiffWithinAt 𝕜 ⊤ (extChartAt I x ∘ (extChartAt I x').symm) (range I) y := - contDiffWithinAt_extend_coord_change I (chart_mem_maximalAtlas I x) (chart_mem_maximalAtlas I x') + contDiffWithinAt_extend_coord_change (chart_mem_maximalAtlas x) (chart_mem_maximalAtlas x') hy +variable (I I') in /-- Conjugating a function to write it in the preferred charts around `x`. The manifold derivative of `f` will just be the derivative of this conjugated function. -/ @[simp, mfld_simps] diff --git a/Mathlib/Geometry/Manifold/VectorBundle/Basic.lean b/Mathlib/Geometry/Manifold/VectorBundle/Basic.lean index 02abb186cc8f9..001e1a6509395 100644 --- a/Mathlib/Geometry/Manifold/VectorBundle/Basic.lean +++ b/Mathlib/Geometry/Manifold/VectorBundle/Basic.lean @@ -122,7 +122,7 @@ section variable [NontriviallyNormedField 𝕜] [NormedAddCommGroup F] [NormedSpace 𝕜 F] [TopologicalSpace (TotalSpace F E)] [∀ x, TopologicalSpace (E x)] {EB : Type*} [NormedAddCommGroup EB] [NormedSpace 𝕜 EB] {HB : Type*} [TopologicalSpace HB] - (IB : ModelWithCorners 𝕜 EB HB) (E' : B → Type*) [∀ x, Zero (E' x)] {EM : Type*} + {IB : ModelWithCorners 𝕜 EB HB} (E' : B → Type*) [∀ x, Zero (E' x)] {EM : Type*} [NormedAddCommGroup EM] [NormedSpace 𝕜 EM] {HM : Type*} [TopologicalSpace HM] {IM : ModelWithCorners 𝕜 EM HM} [TopologicalSpace M] [ChartedSpace HM M] {n : ℕ∞} @@ -149,13 +149,13 @@ theorem FiberBundle.writtenInExtChartAt_trivializationAt {x : TotalSpace F E} {y (hy : y ∈ (extChartAt (IB.prod 𝓘(𝕜, F)) x).target) : writtenInExtChartAt (IB.prod 𝓘(𝕜, F)) (IB.prod 𝓘(𝕜, F)) x (trivializationAt F E x.proj) y = y := - writtenInExtChartAt_chartAt_comp _ _ hy + writtenInExtChartAt_chartAt_comp _ hy theorem FiberBundle.writtenInExtChartAt_trivializationAt_symm {x : TotalSpace F E} {y} (hy : y ∈ (extChartAt (IB.prod 𝓘(𝕜, F)) x).target) : writtenInExtChartAt (IB.prod 𝓘(𝕜, F)) (IB.prod 𝓘(𝕜, F)) (trivializationAt F E x.proj x) (trivializationAt F E x.proj).toPartialHomeomorph.symm y = y := - writtenInExtChartAt_chartAt_symm_comp _ _ hy + writtenInExtChartAt_chartAt_symm_comp _ hy /-! ### Smoothness of maps in/out fiber bundles @@ -165,8 +165,6 @@ bundle at all, just that it is a fiber bundle over a charted base space. namespace Bundle -variable {IB} - /-- Characterization of C^n functions into a smooth vector bundle. -/ theorem contMDiffWithinAt_totalSpace (f : M → TotalSpace F E) {s : Set M} {x₀ : M} : ContMDiffWithinAt IM (IB.prod 𝓘(𝕜, F)) n f s x₀ ↔ @@ -251,7 +249,7 @@ end variable [NontriviallyNormedField 𝕜] {EB : Type*} [NormedAddCommGroup EB] [NormedSpace 𝕜 EB] - {HB : Type*} [TopologicalSpace HB] (IB : ModelWithCorners 𝕜 EB HB) [TopologicalSpace B] + {HB : Type*} [TopologicalSpace HB] {IB : ModelWithCorners 𝕜 EB HB} [TopologicalSpace B] [ChartedSpace HB B] {EM : Type*} [NormedAddCommGroup EM] [NormedSpace 𝕜 EM] {HM : Type*} [TopologicalSpace HM] {IM : ModelWithCorners 𝕜 EM HM} [TopologicalSpace M] [ChartedSpace HM M] {n : ℕ∞} @@ -262,6 +260,7 @@ section WithTopology variable [TopologicalSpace (TotalSpace F E)] [∀ x, TopologicalSpace (E x)] (F E) variable [FiberBundle F E] [VectorBundle 𝕜 F E] +variable (IB) in /-- When `B` is a smooth manifold with corners with respect to a model `IB` and `E` is a topological vector bundle over `B` with fibers isomorphic to `F`, then `SmoothVectorBundle F E IB` registers that the bundle is smooth, in the sense of having smooth transition functions. @@ -294,31 +293,30 @@ theorem smoothOn_symm_coordChangeL : theorem contMDiffOn_coordChangeL : ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) n (fun b : B => (e.coordChangeL 𝕜 e' b : F →L[𝕜] F)) (e.baseSet ∩ e'.baseSet) := - (smoothOn_coordChangeL IB e e').of_le le_top + (smoothOn_coordChangeL e e').of_le le_top theorem contMDiffOn_symm_coordChangeL : ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) n (fun b : B => ((e.coordChangeL 𝕜 e' b).symm : F →L[𝕜] F)) (e.baseSet ∩ e'.baseSet) := - (smoothOn_symm_coordChangeL IB e e').of_le le_top + (smoothOn_symm_coordChangeL e e').of_le le_top variable {e e'} theorem contMDiffAt_coordChangeL {x : B} (h : x ∈ e.baseSet) (h' : x ∈ e'.baseSet) : ContMDiffAt IB 𝓘(𝕜, F →L[𝕜] F) n (fun b : B => (e.coordChangeL 𝕜 e' b : F →L[𝕜] F)) x := - (contMDiffOn_coordChangeL IB e e').contMDiffAt <| + (contMDiffOn_coordChangeL e e').contMDiffAt <| (e.open_baseSet.inter e'.open_baseSet).mem_nhds ⟨h, h'⟩ theorem smoothAt_coordChangeL {x : B} (h : x ∈ e.baseSet) (h' : x ∈ e'.baseSet) : SmoothAt IB 𝓘(𝕜, F →L[𝕜] F) (fun b : B => (e.coordChangeL 𝕜 e' b : F →L[𝕜] F)) x := - contMDiffAt_coordChangeL IB h h' + contMDiffAt_coordChangeL h h' -variable {IB} variable {s : Set M} {f : M → B} {g : M → F} {x : M} protected theorem ContMDiffWithinAt.coordChangeL (hf : ContMDiffWithinAt IM IB n f s x) (he : f x ∈ e.baseSet) (he' : f x ∈ e'.baseSet) : ContMDiffWithinAt IM 𝓘(𝕜, F →L[𝕜] F) n (fun y ↦ (e.coordChangeL 𝕜 e' (f y) : F →L[𝕜] F)) s x := - (contMDiffAt_coordChangeL IB he he').comp_contMDiffWithinAt _ hf + (contMDiffAt_coordChangeL he he').comp_contMDiffWithinAt _ hf protected nonrec theorem ContMDiffAt.coordChangeL (hf : ContMDiffAt IM IB n f x) (he : f x ∈ e.baseSet) (he' : f x ∈ e'.baseSet) : @@ -454,8 +452,8 @@ instance SmoothFiberwiseLinear.hasGroupoid : haveI : MemTrivializationAtlas e := ⟨he⟩ haveI : MemTrivializationAtlas e' := ⟨he'⟩ rw [mem_smoothFiberwiseLinear_iff] - refine ⟨_, _, e.open_baseSet.inter e'.open_baseSet, smoothOn_coordChangeL IB e e', - smoothOn_symm_coordChangeL IB e e', ?_⟩ + refine ⟨_, _, e.open_baseSet.inter e'.open_baseSet, smoothOn_coordChangeL e e', + smoothOn_symm_coordChangeL e e', ?_⟩ refine PartialHomeomorph.eqOnSourceSetoid.symm ⟨?_, ?_⟩ · simp only [e.symm_trans_source_eq e', FiberwiseLinear.partialHomeomorph, trans_toPartialEquiv, symm_toPartialEquiv] @@ -497,7 +495,7 @@ theorem Trivialization.contMDiffAt_iff {f : M → TotalSpace F E} {x₀ : M} (he ContMDiffAt IM (IB.prod 𝓘(𝕜, F)) n f x₀ ↔ ContMDiffAt IM IB n (fun x => (f x).proj) x₀ ∧ ContMDiffAt IM 𝓘(𝕜, F) n (fun x ↦ (e (f x)).2) x₀ := - e.contMDiffWithinAt_iff _ he + e.contMDiffWithinAt_iff he theorem Trivialization.contMDiffOn_iff {f : M → TotalSpace F E} {s : Set M} (he : MapsTo f s e.source) : @@ -505,46 +503,46 @@ theorem Trivialization.contMDiffOn_iff {f : M → TotalSpace F E} {s : Set M} ContMDiffOn IM IB n (fun x => (f x).proj) s ∧ ContMDiffOn IM 𝓘(𝕜, F) n (fun x ↦ (e (f x)).2) s := by simp only [ContMDiffOn, ← forall_and] - exact forall₂_congr fun x hx ↦ e.contMDiffWithinAt_iff IB (he hx) + exact forall₂_congr fun x hx ↦ e.contMDiffWithinAt_iff (he hx) theorem Trivialization.contMDiff_iff {f : M → TotalSpace F E} (he : ∀ x, f x ∈ e.source) : ContMDiff IM (IB.prod 𝓘(𝕜, F)) n f ↔ ContMDiff IM IB n (fun x => (f x).proj) ∧ ContMDiff IM 𝓘(𝕜, F) n (fun x ↦ (e (f x)).2) := - (forall_congr' fun x ↦ e.contMDiffAt_iff IB (he x)).trans forall_and + (forall_congr' fun x ↦ e.contMDiffAt_iff (he x)).trans forall_and theorem Trivialization.smoothWithinAt_iff {f : M → TotalSpace F E} {s : Set M} {x₀ : M} (he : f x₀ ∈ e.source) : SmoothWithinAt IM (IB.prod 𝓘(𝕜, F)) f s x₀ ↔ SmoothWithinAt IM IB (fun x => (f x).proj) s x₀ ∧ SmoothWithinAt IM 𝓘(𝕜, F) (fun x ↦ (e (f x)).2) s x₀ := - e.contMDiffWithinAt_iff IB he + e.contMDiffWithinAt_iff he theorem Trivialization.smoothAt_iff {f : M → TotalSpace F E} {x₀ : M} (he : f x₀ ∈ e.source) : SmoothAt IM (IB.prod 𝓘(𝕜, F)) f x₀ ↔ SmoothAt IM IB (fun x => (f x).proj) x₀ ∧ SmoothAt IM 𝓘(𝕜, F) (fun x ↦ (e (f x)).2) x₀ := - e.contMDiffAt_iff IB he + e.contMDiffAt_iff he theorem Trivialization.smoothOn_iff {f : M → TotalSpace F E} {s : Set M} (he : MapsTo f s e.source) : SmoothOn IM (IB.prod 𝓘(𝕜, F)) f s ↔ SmoothOn IM IB (fun x => (f x).proj) s ∧ SmoothOn IM 𝓘(𝕜, F) (fun x ↦ (e (f x)).2) s := - e.contMDiffOn_iff IB he + e.contMDiffOn_iff he theorem Trivialization.smooth_iff {f : M → TotalSpace F E} (he : ∀ x, f x ∈ e.source) : Smooth IM (IB.prod 𝓘(𝕜, F)) f ↔ Smooth IM IB (fun x => (f x).proj) ∧ Smooth IM 𝓘(𝕜, F) (fun x ↦ (e (f x)).2) := - e.contMDiff_iff IB he + e.contMDiff_iff he theorem Trivialization.smoothOn (e : Trivialization F (π F E)) [MemTrivializationAtlas e] : SmoothOn (IB.prod 𝓘(𝕜, F)) (IB.prod 𝓘(𝕜, F)) e e.source := by have : SmoothOn (IB.prod 𝓘(𝕜, F)) (IB.prod 𝓘(𝕜, F)) id e.source := smoothOn_id - rw [e.smoothOn_iff IB (mapsTo_id _)] at this + rw [e.smoothOn_iff (mapsTo_id _)] at this exact (this.1.prod_mk this.2).congr fun x hx ↦ (e.mk_proj_snd hx).symm theorem Trivialization.smoothOn_symm (e : Trivialization F (π F E)) [MemTrivializationAtlas e] : SmoothOn (IB.prod 𝓘(𝕜, F)) (IB.prod 𝓘(𝕜, F)) e.toPartialHomeomorph.symm e.target := by - rw [e.smoothOn_iff IB e.toPartialHomeomorph.symm_mapsTo] + rw [e.smoothOn_iff e.toPartialHomeomorph.symm_mapsTo] refine ⟨smoothOn_fst.congr fun x hx ↦ e.proj_symm_apply hx, smoothOn_snd.congr fun x hx ↦ ?_⟩ rw [e.apply_symm_apply hx] @@ -614,10 +612,10 @@ instance Bundle.Prod.smoothVectorBundle : SmoothVectorBundle (F₁ × F₂) (E rw [SmoothOn] refine ContMDiffOn.congr ?_ (e₁.coordChangeL_prod 𝕜 e₁' e₂ e₂') refine ContMDiffOn.clm_prodMap ?_ ?_ - · refine (smoothOn_coordChangeL IB e₁ e₁').mono ?_ + · refine (smoothOn_coordChangeL e₁ e₁').mono ?_ simp only [Trivialization.baseSet_prod, mfld_simps] mfld_set_tac - · refine (smoothOn_coordChangeL IB e₂ e₂').mono ?_ + · refine (smoothOn_coordChangeL e₂ e₂').mono ?_ simp only [Trivialization.baseSet_prod, mfld_simps] mfld_set_tac @@ -631,6 +629,7 @@ namespace VectorPrebundle variable [∀ x, TopologicalSpace (E x)] +variable (IB) in /-- Mixin for a `VectorPrebundle` stating smoothness of coordinate changes. -/ class IsSmooth (a : VectorPrebundle 𝕜 F E) : Prop where exists_smoothCoordChange : @@ -642,6 +641,7 @@ class IsSmooth (a : VectorPrebundle 𝕜 F E) : Prop where variable (a : VectorPrebundle 𝕜 F E) [ha : a.IsSmooth IB] {e e' : Pretrivialization F (π F E)} +variable (IB) in /-- A randomly chosen coordinate change on a `SmoothVectorPrebundle`, given by the field `exists_coordChange`. Note that `a.smoothCoordChange` need not be the same as `a.coordChange`. -/ @@ -649,8 +649,6 @@ noncomputable def smoothCoordChange (he : e ∈ a.pretrivializationAtlas) (he' : e' ∈ a.pretrivializationAtlas) (b : B) : F →L[𝕜] F := Classical.choose (ha.exists_smoothCoordChange e he e' he') b -variable {IB} - theorem smoothOn_smoothCoordChange (he : e ∈ a.pretrivializationAtlas) (he' : e' ∈ a.pretrivializationAtlas) : SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (a.smoothCoordChange IB he he') (e.baseSet ∩ e'.baseSet) := @@ -669,7 +667,7 @@ theorem mk_smoothCoordChange (he : e ∈ a.pretrivializationAtlas) rw [e.proj_symm_apply' hb.1]; exact hb.2 · exact a.smoothCoordChange_apply he he' hb v -variable (IB) +variable (IB) in /-- Make a `SmoothVectorBundle` from a `SmoothVectorPrebundle`. -/ theorem smoothVectorBundle : @SmoothVectorBundle _ _ F E _ _ _ _ _ _ IB _ _ _ _ _ _ a.totalSpaceTopology _ a.toFiberBundle a.toVectorBundle := diff --git a/Mathlib/Geometry/Manifold/VectorBundle/Hom.lean b/Mathlib/Geometry/Manifold/VectorBundle/Hom.lean index 33808cd26cd8e..6708e0a288ed6 100644 --- a/Mathlib/Geometry/Manifold/VectorBundle/Hom.lean +++ b/Mathlib/Geometry/Manifold/VectorBundle/Hom.lean @@ -29,7 +29,7 @@ variable {𝕜 B F₁ F₂ M : Type*} {E₁ : B → Type*} {E₂ : B → Type*} [TopologicalSpace (TotalSpace F₂ E₂)] [∀ x, TopologicalSpace (E₂ x)] {EB : Type*} [NormedAddCommGroup EB] [NormedSpace 𝕜 EB] {HB : Type*} [TopologicalSpace HB] - (IB : ModelWithCorners 𝕜 EB HB) [TopologicalSpace B] [ChartedSpace HB B] {EM : Type*} + {IB : ModelWithCorners 𝕜 EB HB} [TopologicalSpace B] [ChartedSpace HB B] {EM : Type*} [NormedAddCommGroup EM] [NormedSpace 𝕜 EM] {HM : Type*} [TopologicalSpace HM] {IM : ModelWithCorners 𝕜 EM HM} [TopologicalSpace M] [ChartedSpace HM M] {n : ℕ∞} [FiberBundle F₁ E₁] [VectorBundle 𝕜 F₁ E₁] @@ -45,8 +45,8 @@ theorem smoothOn_continuousLinearMapCoordChange SmoothOn IB 𝓘(𝕜, (F₁ →L[𝕜] F₂) →L[𝕜] F₁ →L[𝕜] F₂) (continuousLinearMapCoordChange (RingHom.id 𝕜) e₁ e₁' e₂ e₂') (e₁.baseSet ∩ e₂.baseSet ∩ (e₁'.baseSet ∩ e₂'.baseSet)) := by - have h₁ := smoothOn_coordChangeL IB e₁' e₁ - have h₂ := smoothOn_coordChangeL IB e₂ e₂' + have h₁ := smoothOn_coordChangeL (IB := IB) e₁' e₁ + have h₂ := smoothOn_coordChangeL (IB := IB) e₂ e₂' refine (h₁.mono ?_).cle_arrowCongr (h₂.mono ?_) <;> mfld_set_tac variable [∀ x, TopologicalAddGroup (E₂ x)] [∀ x, ContinuousSMul 𝕜 (E₂ x)] @@ -58,8 +58,6 @@ theorem hom_chart (y₀ y : LE₁E₂) : Trivialization.coe_coe, PartialHomeomorph.refl_apply, Function.id_def, hom_trivializationAt_apply] -variable {IB} - theorem contMDiffAt_hom_bundle (f : M → LE₁E₂) {x₀ : M} {n : ℕ∞} : ContMDiffAt IM (IB.prod 𝓘(𝕜, F₁ →L[𝕜] F₂)) n f x₀ ↔ ContMDiffAt IM IB n (fun x => (f x).1) x₀ ∧ @@ -81,7 +79,7 @@ instance Bundle.ContinuousLinearMap.vectorPrebundle.isSmooth : exists_smoothCoordChange := by rintro _ ⟨e₁, e₂, he₁, he₂, rfl⟩ _ ⟨e₁', e₂', he₁', he₂', rfl⟩ exact ⟨continuousLinearMapCoordChange (RingHom.id 𝕜) e₁ e₁' e₂ e₂', - smoothOn_continuousLinearMapCoordChange IB, + smoothOn_continuousLinearMapCoordChange, continuousLinearMapCoordChange_apply (RingHom.id 𝕜) e₁ e₁' e₂ e₂'⟩ instance SmoothVectorBundle.continuousLinearMap : diff --git a/Mathlib/Geometry/Manifold/VectorBundle/Pullback.lean b/Mathlib/Geometry/Manifold/VectorBundle/Pullback.lean index 30454293a1bd9..3671bb2677dac 100644 --- a/Mathlib/Geometry/Manifold/VectorBundle/Pullback.lean +++ b/Mathlib/Geometry/Manifold/VectorBundle/Pullback.lean @@ -24,7 +24,7 @@ variable {𝕜 B B' M : Type*} (F : Type*) (E : B → Type*) variable [NontriviallyNormedField 𝕜] [∀ x, AddCommMonoid (E x)] [∀ x, Module 𝕜 (E x)] [NormedAddCommGroup F] [NormedSpace 𝕜 F] [TopologicalSpace (TotalSpace F E)] [∀ x, TopologicalSpace (E x)] {EB : Type*} [NormedAddCommGroup EB] [NormedSpace 𝕜 EB] - {HB : Type*} [TopologicalSpace HB] (IB : ModelWithCorners 𝕜 EB HB) [TopologicalSpace B] + {HB : Type*} [TopologicalSpace HB] {IB : ModelWithCorners 𝕜 EB HB} [TopologicalSpace B] [ChartedSpace HB B] [SmoothManifoldWithCorners IB B] {EB' : Type*} [NormedAddCommGroup EB'] [NormedSpace 𝕜 EB'] {HB' : Type*} [TopologicalSpace HB'] (IB' : ModelWithCorners 𝕜 EB' HB') [TopologicalSpace B'] [ChartedSpace HB' B'] [SmoothManifoldWithCorners IB' B'] [FiberBundle F E] @@ -35,7 +35,7 @@ vector bundle `f *ᵖ E` is a smooth vector bundle. -/ instance SmoothVectorBundle.pullback : SmoothVectorBundle F (f *ᵖ E) IB' where smoothOn_coordChangeL := by rintro _ _ ⟨e, he, rfl⟩ ⟨e', he', rfl⟩ - refine ((smoothOn_coordChangeL _ e e').comp f.smooth.smoothOn fun b hb => hb).congr ?_ + refine ((smoothOn_coordChangeL e e').comp f.smooth.smoothOn fun b hb => hb).congr ?_ rintro b (hb : f b ∈ e.baseSet ∩ e'.baseSet); ext v show ((e.pullback f).coordChangeL 𝕜 (e'.pullback f) b) v = (e.coordChangeL 𝕜 e' (f b)) v rw [e.coordChangeL_apply e' hb, (e.pullback f).coordChangeL_apply' _] diff --git a/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean b/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean index 7c85fe03db9fd..b7ae7412cc9d6 100644 --- a/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean +++ b/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean @@ -49,7 +49,6 @@ variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedAddCom [SmoothManifoldWithCorners I M] {M' : Type*} [TopologicalSpace M'] [ChartedSpace H' M'] [SmoothManifoldWithCorners I' M'] {F : Type*} [NormedAddCommGroup F] [NormedSpace 𝕜 F] -variable (I) /-- Auxiliary lemma for tangent spaces: the derivative of a coordinate change between two charts is smooth on its source. -/ @@ -60,14 +59,14 @@ theorem contDiffOn_fderiv_coord_change (i j : atlas H M) : rw [i.1.extend_coord_change_source]; apply image_subset_range intro x hx refine (ContDiffWithinAt.fderivWithin_right ?_ I.uniqueDiffOn le_top <| h hx).mono h - refine (PartialHomeomorph.contDiffOn_extend_coord_change I (subset_maximalAtlas I j.2) - (subset_maximalAtlas I i.2) x hx).mono_of_mem ?_ - exact i.1.extend_coord_change_source_mem_nhdsWithin j.1 I hx + refine (PartialHomeomorph.contDiffOn_extend_coord_change (subset_maximalAtlas j.2) + (subset_maximalAtlas i.2) x hx).mono_of_mem ?_ + exact i.1.extend_coord_change_source_mem_nhdsWithin j.1 hx -variable (M) open SmoothManifoldWithCorners +variable (I M) in /-- Let `M` be a smooth manifold with corners with model `I` on `(E, H)`. Then `tangentBundleCore I M` is the vector bundle core for the tangent bundle over `M`. It is indexed by the atlas of `M`, with fiber `E` and its change of coordinates from the chart `i` @@ -89,35 +88,33 @@ def tangentBundleCore : VectorBundleCore 𝕜 M E (atlas H M) where simp only rw [Filter.EventuallyEq.fderivWithin_eq, fderivWithin_id', ContinuousLinearMap.id_apply] · exact I.uniqueDiffWithinAt_image - · filter_upwards [i.1.extend_target_mem_nhdsWithin I hx] with y hy + · filter_upwards [i.1.extend_target_mem_nhdsWithin hx] with y hy exact (i.1.extend I).right_inv hy - · simp_rw [Function.comp_apply, i.1.extend_left_inv I hx] + · simp_rw [Function.comp_apply, i.1.extend_left_inv hx] continuousOn_coordChange i j := by - refine (contDiffOn_fderiv_coord_change I i j).continuousOn.comp - ((i.1.continuousOn_extend I).mono ?_) ?_ + refine (contDiffOn_fderiv_coord_change i j).continuousOn.comp + (i.1.continuousOn_extend.mono ?_) ?_ · rw [i.1.extend_source]; exact inter_subset_left simp_rw [← i.1.extend_image_source_inter, mapsTo_image] coordChange_comp := by rintro i j k x ⟨⟨hxi, hxj⟩, hxk⟩ v rw [fderivWithin_fderivWithin, Filter.EventuallyEq.fderivWithin_eq] - · have := i.1.extend_preimage_mem_nhds I hxi (j.1.extend_source_mem_nhds I hxj) + · have := i.1.extend_preimage_mem_nhds (I := I) hxi (j.1.extend_source_mem_nhds (I := I) hxj) filter_upwards [nhdsWithin_le_nhds this] with y hy simp_rw [Function.comp_apply, (j.1.extend I).left_inv hy] - · simp_rw [Function.comp_apply, i.1.extend_left_inv I hxi, j.1.extend_left_inv I hxj] - · exact (contDiffWithinAt_extend_coord_change' I (subset_maximalAtlas I k.2) - (subset_maximalAtlas I j.2) hxk hxj).differentiableWithinAt le_top - · exact (contDiffWithinAt_extend_coord_change' I (subset_maximalAtlas I j.2) - (subset_maximalAtlas I i.2) hxj hxi).differentiableWithinAt le_top + · simp_rw [Function.comp_apply, i.1.extend_left_inv hxi, j.1.extend_left_inv hxj] + · exact (contDiffWithinAt_extend_coord_change' (subset_maximalAtlas k.2) + (subset_maximalAtlas j.2) hxk hxj).differentiableWithinAt le_top + · exact (contDiffWithinAt_extend_coord_change' (subset_maximalAtlas j.2) + (subset_maximalAtlas i.2) hxj hxi).differentiableWithinAt le_top · intro x _; exact mem_range_self _ · exact I.uniqueDiffWithinAt_image - · rw [Function.comp_apply, i.1.extend_left_inv I hxi] + · rw [Function.comp_apply, i.1.extend_left_inv hxi] -- Porting note: moved to a separate `simp high` lemma b/c `simp` can simplify the LHS @[simp high] theorem tangentBundleCore_baseSet (i) : (tangentBundleCore I M).baseSet i = i.1.source := rfl -variable {M} - theorem tangentBundleCore_coordChange_achart (x x' z : M) : (tangentBundleCore I M).coordChange (achart H x) (achart H x') z = fderivWithin 𝕜 (extChartAt I x' ∘ (extChartAt I x).symm) (range I) (extChartAt I x z) := @@ -125,6 +122,7 @@ theorem tangentBundleCore_coordChange_achart (x x' z : M) : section tangentCoordChange +variable (I) in /-- In a manifold `M`, given two preferred charts indexed by `x y : M`, `tangentCoordChange I x y` is the family of derivatives of the corresponding change-of-coordinates map. It takes junk values outside the intersection of the sources of the two charts. @@ -134,8 +132,6 @@ sets as the preferred charts of the base manifold. -/ abbrev tangentCoordChange (x y : M) : M → E →L[𝕜] E := (tangentBundleCore I M).coordChange (achart H x) (achart H y) -variable {I} - lemma tangentCoordChange_def {x y z : M} : tangentCoordChange I x y z = fderivWithin 𝕜 (extChartAt I y ∘ (extChartAt I x).symm) (range I) (extChartAt I x z) := rfl @@ -159,7 +155,7 @@ lemma hasFDerivWithinAt_tangentCoordChange {x y z : M} have h' : extChartAt I x z ∈ ((extChartAt I x).symm ≫ (extChartAt I y)).source := by rw [PartialEquiv.trans_source'', PartialEquiv.symm_symm, PartialEquiv.symm_target] exact mem_image_of_mem _ h - ((contDiffWithinAt_ext_coord_change I y x h').differentiableWithinAt (by simp)).hasFDerivWithinAt + ((contDiffWithinAt_ext_coord_change y x h').differentiableWithinAt (by simp)).hasFDerivWithinAt lemma continuousOn_tangentCoordChange (x y : M) : ContinuousOn (tangentCoordChange I x y) ((extChartAt I x).source ∩ (extChartAt I y).source) := by @@ -168,8 +164,6 @@ lemma continuousOn_tangentCoordChange (x y : M) : ContinuousOn (tangentCoordChan end tangentCoordChange -variable (M) - local notation "TM" => TangentBundle I M section TangentBundleInstances @@ -292,16 +286,16 @@ end TangentBundle instance tangentBundleCore.isSmooth : (tangentBundleCore I M).IsSmooth I := by refine ⟨fun i j => ?_⟩ - rw [SmoothOn, contMDiffOn_iff_source_of_mem_maximalAtlas (subset_maximalAtlas I i.2), + rw [SmoothOn, contMDiffOn_iff_source_of_mem_maximalAtlas (subset_maximalAtlas i.2), contMDiffOn_iff_contDiffOn] - · refine ((contDiffOn_fderiv_coord_change I i j).congr fun x hx => ?_).mono ?_ + · refine ((contDiffOn_fderiv_coord_change (I := I) i j).congr fun x hx => ?_).mono ?_ · rw [PartialEquiv.trans_source'] at hx simp_rw [Function.comp_apply, tangentBundleCore_coordChange, (i.1.extend I).right_inv hx.1] - · exact (i.1.extend_image_source_inter j.1 I).subset + · exact (i.1.extend_image_source_inter j.1).subset · apply inter_subset_left instance TangentBundle.smoothVectorBundle : SmoothVectorBundle E (TangentSpace I : M → Type _) I := - (tangentBundleCore I M).smoothVectorBundle _ + (tangentBundleCore I M).smoothVectorBundle end TangentBundleInstances @@ -341,8 +335,7 @@ theorem tangentBundleCore_coordChange_model_space (x x' z : H) : ContinuousLinearMap.id 𝕜 E := by ext v; exact (tangentBundleCore I H).coordChange_self (achart _ z) z (mem_univ _) v -variable (H) - +variable (I) in /-- The canonical identification between the tangent bundle to the model space and the product, as a homeomorphism -/ def tangentBundleModelSpaceHomeomorph : TangentBundle I H ≃ₜ ModelProd H E := @@ -364,19 +357,18 @@ def tangentBundleModelSpaceHomeomorph : TangentBundle I H ≃ₜ ModelProd H E : @[simp, mfld_simps] theorem tangentBundleModelSpaceHomeomorph_coe : - (tangentBundleModelSpaceHomeomorph H I : TangentBundle I H → ModelProd H E) = + (tangentBundleModelSpaceHomeomorph I : TangentBundle I H → ModelProd H E) = TotalSpace.toProd H E := rfl @[simp, mfld_simps] theorem tangentBundleModelSpaceHomeomorph_coe_symm : - ((tangentBundleModelSpaceHomeomorph H I).symm : ModelProd H E → TangentBundle I H) = + ((tangentBundleModelSpaceHomeomorph I).symm : ModelProd H E → TangentBundle I H) = (TotalSpace.toProd H E).symm := rfl section inTangentCoordinates -variable (I') {M H} variable {N : Type*} /-- The map `in_coordinates` for the tangent bundle is trivial on the model spaces -/ @@ -386,6 +378,7 @@ theorem inCoordinates_tangent_bundle_core_model_space (x₀ x : H) (y₀ y : H') simp_rw [tangentBundleCore_indexAt, tangentBundleCore_coordChange_model_space, ContinuousLinearMap.id_comp, ContinuousLinearMap.comp_id] +variable (I I') in /-- When `ϕ x` is a continuous linear map that changes vectors in charts around `f x` to vectors in charts around `g x`, `inTangentCoordinates I I' f g ϕ x₀ x` is a coordinate change of this continuous linear map that makes sense from charts around `f x₀` to charts around `g x₀` diff --git a/Mathlib/Geometry/Manifold/WhitneyEmbedding.lean b/Mathlib/Geometry/Manifold/WhitneyEmbedding.lean index 4e56b5f717089..8106c8dd8e80c 100644 --- a/Mathlib/Geometry/Manifold/WhitneyEmbedding.lean +++ b/Mathlib/Geometry/Manifold/WhitneyEmbedding.lean @@ -83,7 +83,7 @@ theorem comp_embeddingPiTangent_mfderiv (x : M) (hx : x ∈ s) : (@ContinuousLinearMap.proj ℝ _ ι (fun _ => E × ℝ) _ _ (fun _ => inferInstance) (f.ind x hx)) have := L.hasMFDerivAt.comp x f.embeddingPiTangent.smooth.mdifferentiableAt.hasMFDerivAt convert hasMFDerivAt_unique this _ - refine (hasMFDerivAt_extChartAt I (f.mem_chartAt_ind_source x hx)).congr_of_eventuallyEq ?_ + refine (hasMFDerivAt_extChartAt (f.mem_chartAt_ind_source x hx)).congr_of_eventuallyEq ?_ refine (f.eventuallyEq_one x hx).mono fun y hy => ?_ simp only [L, embeddingPiTangent_coe, ContinuousLinearMap.coe_comp', (· ∘ ·), ContinuousLinearMap.coe_fst', ContinuousLinearMap.proj_apply] @@ -92,7 +92,7 @@ theorem comp_embeddingPiTangent_mfderiv (x : M) (hx : x ∈ s) : theorem embeddingPiTangent_ker_mfderiv (x : M) (hx : x ∈ s) : LinearMap.ker (mfderiv I 𝓘(ℝ, ι → E × ℝ) f.embeddingPiTangent x) = ⊥ := by apply bot_unique - rw [← (mdifferentiable_chart I (f.c (f.ind x hx))).ker_mfderiv_eq_bot + rw [← (mdifferentiable_chart (f.c (f.ind x hx))).ker_mfderiv_eq_bot (f.mem_chartAt_ind_source x hx), ← comp_embeddingPiTangent_mfderiv] exact LinearMap.ker_le_ker_comp _ _ diff --git a/Mathlib/NumberTheory/ModularForms/Basic.lean b/Mathlib/NumberTheory/ModularForms/Basic.lean index aca96ca362386..c0482b00817f7 100644 --- a/Mathlib/NumberTheory/ModularForms/Basic.lean +++ b/Mathlib/NumberTheory/ModularForms/Basic.lean @@ -152,7 +152,7 @@ theorem add_apply (f g : ModularForm Γ k) (z : ℍ) : (f + g) z = f z + g z := instance instZero : Zero (ModularForm Γ k) := ⟨ { toSlashInvariantForm := 0 - holo' := fun _ => mdifferentiableAt_const 𝓘(ℂ, ℂ) 𝓘(ℂ, ℂ) + holo' := fun _ => mdifferentiableAt_const bdd_at_infty' := fun A => by simpa using zero_form_isBoundedAtImInfty } ⟩ @[simp] @@ -245,7 +245,7 @@ theorem mul_coe {k_1 k_2 : ℤ} {Γ : Subgroup SL(2, ℤ)} (f : ModularForm Γ k @[simps! (config := .asFn) toFun toSlashInvariantForm] def const (x : ℂ) : ModularForm Γ 0 where toSlashInvariantForm := .const x - holo' _ := mdifferentiableAt_const 𝓘(ℂ, ℂ) 𝓘(ℂ, ℂ) + holo' _ := mdifferentiableAt_const bdd_at_infty' A := by simpa only [SlashInvariantForm.const_toFun, ModularForm.is_invariant_const] using atImInfty.const_boundedAtFilter x @@ -301,7 +301,7 @@ theorem add_apply (f g : CuspForm Γ k) (z : ℍ) : (f + g) z = f z + g z := instance instZero : Zero (CuspForm Γ k) := ⟨ { toSlashInvariantForm := 0 - holo' := fun _ => mdifferentiableAt_const 𝓘(ℂ, ℂ) 𝓘(ℂ, ℂ) + holo' := fun _ => mdifferentiableAt_const zero_at_infty' := by simpa using Filter.zero_zeroAtFilter _ } ⟩ @[simp] From 9d8ba939055398dc4f21d40df61fe8419bafe164 Mon Sep 17 00:00:00 2001 From: Hannah Scholz Date: Thu, 24 Oct 2024 19:28:37 +0000 Subject: [PATCH 392/505] feat: `T2Space` is `Antitone` (#18188) --- Mathlib/Topology/Separation.lean | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/Mathlib/Topology/Separation.lean b/Mathlib/Topology/Separation.lean index a99e1b615e460..284f7ef7efa3c 100644 --- a/Mathlib/Topology/Separation.lean +++ b/Mathlib/Topology/Separation.lean @@ -1892,6 +1892,10 @@ theorem not_preirreducible_nontrivial_t2 (X) [TopologicalSpace X] [Preirreducibl [Nontrivial X] [T2Space X] : False := (PreirreducibleSpace.isPreirreducible_univ (X := X)).subsingleton.not_nontrivial nontrivial_univ +theorem t2Space_antitone {X : Type*} : Antitone (@T2Space X) := + fun inst₁ inst₂ h_top h_t2 ↦ @T2Space.of_injective_continuous _ _ inst₁ inst₂ + h_t2 _ Function.injective_id <| continuous_id_of_le h_top + end Separation section RegularSpace From 1cc8fd2224b0e1f1a0c7ab16628f98fc404e837a Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Thu, 24 Oct 2024 19:49:38 +0000 Subject: [PATCH 393/505] chore(Algebra/Group/Action): reduce theory included with `Group.Action.Defs` (#18187) The `Group/Action/Defs.lean` file depends on a bit of theory and in turn defines some results not needed for e.g. modules. This PR tries to ensure that the focus is on the definitions of `MulAction`, `SMulCommClass` and `IsScalarTower`. In the process, I made the following moves: * `Action/End.lean` now contains the definition of `Function.End` and some results relating homomorphisms with multiplication * `Action/Faithful.lean` now contains the definition of `FaithfulAction` * `Action/Pretransitive.lean` now contains the definition of `PretransitiveAction` * `Action/TypeTags.lean` now contains the `Additive`/`Multiplicative` instances. --- Mathlib.lean | 4 + Mathlib/Algebra/Group/Action/Defs.lean | 291 +----------------- Mathlib/Algebra/Group/Action/End.lean | 169 ++++++++++ Mathlib/Algebra/Group/Action/Faithful.lean | 61 ++++ Mathlib/Algebra/Group/Action/Opposite.lean | 3 +- Mathlib/Algebra/Group/Action/Option.lean | 2 +- Mathlib/Algebra/Group/Action/Pi.lean | 2 +- .../Algebra/Group/Action/Pretransitive.lean | 130 ++++++++ Mathlib/Algebra/Group/Action/Prod.lean | 2 +- Mathlib/Algebra/Group/Action/Sigma.lean | 2 +- Mathlib/Algebra/Group/Action/Sum.lean | 2 +- Mathlib/Algebra/Group/Action/TypeTags.lean | 72 +++++ Mathlib/Algebra/Group/Action/Units.lean | 2 +- .../Algebra/Group/Submonoid/Operations.lean | 2 +- 14 files changed, 450 insertions(+), 294 deletions(-) create mode 100644 Mathlib/Algebra/Group/Action/End.lean create mode 100644 Mathlib/Algebra/Group/Action/Faithful.lean create mode 100644 Mathlib/Algebra/Group/Action/Pretransitive.lean create mode 100644 Mathlib/Algebra/Group/Action/TypeTags.lean diff --git a/Mathlib.lean b/Mathlib.lean index 5c1737f14e072..22dc00ac223e7 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -227,12 +227,16 @@ import Mathlib.Algebra.GradedMonoid import Mathlib.Algebra.GradedMulAction import Mathlib.Algebra.Group.Action.Basic import Mathlib.Algebra.Group.Action.Defs +import Mathlib.Algebra.Group.Action.End +import Mathlib.Algebra.Group.Action.Faithful import Mathlib.Algebra.Group.Action.Opposite import Mathlib.Algebra.Group.Action.Option import Mathlib.Algebra.Group.Action.Pi +import Mathlib.Algebra.Group.Action.Pretransitive import Mathlib.Algebra.Group.Action.Prod import Mathlib.Algebra.Group.Action.Sigma import Mathlib.Algebra.Group.Action.Sum +import Mathlib.Algebra.Group.Action.TypeTags import Mathlib.Algebra.Group.Action.Units import Mathlib.Algebra.Group.AddChar import Mathlib.Algebra.Group.Aut diff --git a/Mathlib/Algebra/Group/Action/Defs.lean b/Mathlib/Algebra/Group/Action/Defs.lean index ea9c58c15f971..823dbf69f3586 100644 --- a/Mathlib/Algebra/Group/Action/Defs.lean +++ b/Mathlib/Algebra/Group/Action/Defs.lean @@ -4,9 +4,12 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Yury Kudryashov -/ import Mathlib.Algebra.Group.Commute.Defs -import Mathlib.Algebra.Group.TypeTags import Mathlib.Algebra.Opposites +import Mathlib.Data.FunLike.Basic import Mathlib.Logic.Embedding.Basic +import Mathlib.Logic.Function.Iterate +import Mathlib.Logic.Nontrivial.Basic +import Mathlib.Tactic.Spread /-! # Definitions of group actions @@ -22,8 +25,7 @@ notation classes `SMul` and its additive version `VAdd`: The hierarchy is extended further by `Module`, defined elsewhere. -Also provided are typeclasses for faithful and transitive actions, and typeclasses regarding the -interaction of different group actions, +Also provided are typeclasses regarding the interaction of different group actions, * `SMulCommClass M N α` and its additive version `VAddCommClass M N α`; * `IsScalarTower M N α` and its additive version `VAddAssocClass M N α`; @@ -50,26 +52,6 @@ open Function (Injective Surjective) variable {M N G H α β γ δ : Type*} -/-! ### Faithful actions -/ - -/-- Typeclass for faithful actions. -/ -class FaithfulVAdd (G : Type*) (P : Type*) [VAdd G P] : Prop where - /-- Two elements `g₁` and `g₂` are equal whenever they act in the same way on all points. -/ - eq_of_vadd_eq_vadd : ∀ {g₁ g₂ : G}, (∀ p : P, g₁ +ᵥ p = g₂ +ᵥ p) → g₁ = g₂ - -/-- Typeclass for faithful actions. -/ -@[to_additive] -class FaithfulSMul (M : Type*) (α : Type*) [SMul M α] : Prop where - /-- Two elements `m₁` and `m₂` are equal whenever they act in the same way on all points. -/ - eq_of_smul_eq_smul : ∀ {m₁ m₂ : M}, (∀ a : α, m₁ • a = m₂ • a) → m₁ = m₂ - -export FaithfulSMul (eq_of_smul_eq_smul) -export FaithfulVAdd (eq_of_vadd_eq_vadd) - -@[to_additive] -lemma smul_left_injective' [SMul M α] [FaithfulSMul M α] : Injective ((· • ·) : M → α → α) := - fun _ _ h ↦ FaithfulSMul.eq_of_smul_eq_smul (congr_fun h) - -- see Note [lower instance priority] /-- See also `Monoid.toMulAction` and `MulZeroClass.toSMulWithZero`. -/ @[to_additive "See also `AddMonoid.toAddAction`"] @@ -78,11 +60,6 @@ instance (priority := 910) Mul.toSMul (α : Type*) [Mul α] : SMul α α := ⟨( @[to_additive (attr := simp)] lemma smul_eq_mul (α : Type*) [Mul α] {a a' : α} : a • a' = a * a' := rfl -/-- `Monoid.toMulAction` is faithful on cancellative monoids. -/ -@[to_additive " `AddMonoid.toAddAction` is faithful on additive cancellative monoids. "] -instance RightCancelMonoid.faithfulSMul [RightCancelMonoid α] : FaithfulSMul α α := - ⟨fun h ↦ mul_right_cancel (h 1)⟩ - /-- Type class for additive monoid actions. -/ class AddAction (G : Type*) (P : Type*) [AddMonoid G] extends VAdd G P where /-- Zero is a neutral element for `+ᵥ` -/ @@ -98,48 +75,6 @@ class MulAction (α : Type*) (β : Type*) [Monoid α] extends SMul α β where /-- Associativity of `•` and `*` -/ mul_smul : ∀ (x y : α) (b : β), (x * y) • b = x • y • b -/-! -### (Pre)transitive action - -`M` acts pretransitively on `α` if for any `x y` there is `g` such that `g • x = y` (or `g +ᵥ x = y` -for an additive action). A transitive action should furthermore have `α` nonempty. - -In this section we define typeclasses `MulAction.IsPretransitive` and -`AddAction.IsPretransitive` and provide `MulAction.exists_smul_eq`/`AddAction.exists_vadd_eq`, -`MulAction.surjective_smul`/`AddAction.surjective_vadd` as public interface to access this -property. We do not provide typeclasses `*Action.IsTransitive`; users should assume -`[MulAction.IsPretransitive M α] [Nonempty α]` instead. --/ - -/-- `M` acts pretransitively on `α` if for any `x y` there is `g` such that `g +ᵥ x = y`. - A transitive action should furthermore have `α` nonempty. -/ -class AddAction.IsPretransitive (M α : Type*) [VAdd M α] : Prop where - /-- There is `g` such that `g +ᵥ x = y`. -/ - exists_vadd_eq : ∀ x y : α, ∃ g : M, g +ᵥ x = y - -/-- `M` acts pretransitively on `α` if for any `x y` there is `g` such that `g • x = y`. - A transitive action should furthermore have `α` nonempty. -/ -@[to_additive] -class MulAction.IsPretransitive (M α : Type*) [SMul M α] : Prop where - /-- There is `g` such that `g • x = y`. -/ - exists_smul_eq : ∀ x y : α, ∃ g : M, g • x = y - -namespace MulAction -variable (M) [SMul M α] [IsPretransitive M α] - -@[to_additive] -lemma exists_smul_eq (x y : α) : ∃ m : M, m • x = y := IsPretransitive.exists_smul_eq x y - -@[to_additive] -lemma surjective_smul (x : α) : Surjective fun c : M ↦ c • x := exists_smul_eq M x - -/-- The regular action of a group on itself is transitive. -/ -@[to_additive "The regular action of a group on itself is transitive."] -instance Regular.isPretransitive [Group G] : IsPretransitive G G := - ⟨fun x y ↦ ⟨y * x⁻¹, inv_mul_cancel_right _ _⟩⟩ - -end MulAction - /-! ### Scalar tower and commuting actions -/ /-- A typeclass mixin saying that two additive actions on the same space commute. -/ @@ -427,19 +362,6 @@ protected abbrev Function.Surjective.mulAction [SMul M β] (f : α → β) (hf : one_smul := by simp [hf.forall, ← smul] mul_smul := by simp [hf.forall, ← smul, mul_smul] -/-- Push forward the action of `R` on `M` along a compatible surjective map `f : R →* S`. - -See also `Function.Surjective.distribMulActionLeft` and `Function.Surjective.moduleLeft`. --/ -@[to_additive -"Push forward the action of `R` on `M` along a compatible surjective map `f : R →+ S`."] -abbrev Function.Surjective.mulActionLeft {R S M : Type*} [Monoid R] [MulAction R M] [Monoid S] - [SMul S M] (f : R →* S) (hf : Surjective f) (hsmul : ∀ (c) (x : M), f c • x = c • x) : - MulAction S M where - smul := (· • ·) - one_smul b := by rw [← f.map_one, hsmul, one_smul] - mul_smul := hf.forall₂.mpr fun a b x ↦ by simp only [← f.map_mul, hsmul, mul_smul] - section variable (M) @@ -528,55 +450,6 @@ def toFun : α ↪ M → α := @[to_additive (attr := simp)] lemma toFun_apply (x : M) (y : α) : MulAction.toFun M α y x = x • y := rfl -variable (α) - -/-- A multiplicative action of `M` on `α` and a monoid homomorphism `N → M` induce -a multiplicative action of `N` on `α`. - -See note [reducible non-instances]. -/ -@[to_additive] -abbrev compHom [Monoid N] (g : N →* M) : MulAction N α where - smul := SMul.comp.smul g - -- Porting note: was `by simp [g.map_one, MulAction.one_smul]` - one_smul _ := by simpa [(· • ·)] using MulAction.one_smul .. - -- Porting note: was `by simp [g.map_mul, MulAction.mul_smul]` - mul_smul _ _ _ := by simpa [(· • ·)] using MulAction.mul_smul .. - -/-- An additive action of `M` on `α` and an additive monoid homomorphism `N → M` induce -an additive action of `N` on `α`. - -See note [reducible non-instances]. -/ -add_decl_doc AddAction.compHom - -@[to_additive] -lemma compHom_smul_def - {E F G : Type*} [Monoid E] [Monoid F] [MulAction F G] (f : E →* F) (a : E) (x : G) : - letI : MulAction E G := MulAction.compHom _ f - a • x = (f a) • x := rfl - -/-- If an action is transitive, then composing this action with a surjective homomorphism gives -again a transitive action. -/ -@[to_additive] -lemma isPretransitive_compHom {E F G : Type*} [Monoid E] [Monoid F] [MulAction F G] - [IsPretransitive F G] {f : E →* F} (hf : Surjective f) : - letI : MulAction E G := MulAction.compHom _ f - IsPretransitive E G := by - let _ : MulAction E G := MulAction.compHom _ f - refine ⟨fun x y ↦ ?_⟩ - obtain ⟨m, rfl⟩ : ∃ m : F, m • x = y := exists_smul_eq F x y - obtain ⟨e, rfl⟩ : ∃ e, f e = m := hf m - exact ⟨e, rfl⟩ - -@[to_additive] -lemma IsPretransitive.of_smul_eq {M N α : Type*} [SMul M α] [SMul N α] [IsPretransitive M α] - (f : M → N) (hf : ∀ {c : M} {x : α}, f c • x = c • x) : IsPretransitive N α where - exists_smul_eq x y := (exists_smul_eq x y).elim fun m h ↦ ⟨f m, hf.trans h⟩ - -@[to_additive] -lemma IsPretransitive.of_compHom {M N α : Type*} [Monoid M] [Monoid N] [MulAction N α] - (f : M →* N) [h : letI := compHom α f; IsPretransitive M α] : IsPretransitive N α := - letI := compHom α f; h.of_smul_eq f rfl - end MulAction end @@ -587,11 +460,6 @@ lemma smul_one_smul {M} (N) [Monoid N] [SMul M N] [MulAction N α] [SMul M α] [IsScalarTower M N α] (x : M) (y : α) : (x • (1 : N)) • y = x • y := by rw [smul_assoc, one_smul] -@[to_additive] -lemma MulAction.IsPretransitive.of_isScalarTower (M : Type*) {N α : Type*} [Monoid N] [SMul M N] - [MulAction N α] [SMul M α] [IsScalarTower M N α] [IsPretransitive M α] : IsPretransitive N α := - of_smul_eq (fun x : M ↦ x • 1) (smul_one_smul N _ _) - @[to_additive (attr := simp)] lemma smul_one_mul {M N} [MulOneClass N] [SMul M N] [IsScalarTower M N N] (x : M) (y : N) : x • (1 : N) * y = x • y := by rw [smul_mul_assoc, one_mul] @@ -610,153 +478,4 @@ lemma SMulCommClass.of_mul_smul_one {M N} [Monoid N] [SMul M N] (H : ∀ (x : M) (y : N), y * x • (1 : N) = x • y) : SMulCommClass M N N := ⟨fun x y z ↦ by rw [← H x z, smul_eq_mul, ← H, smul_eq_mul, mul_assoc]⟩ -/-- If the multiplicative action of `M` on `N` is compatible with multiplication on `N`, then -`fun x ↦ x • 1` is a monoid homomorphism from `M` to `N`. -/ -@[to_additive (attr := simps) -"If the additive action of `M` on `N` is compatible with addition on `N`, then -`fun x ↦ x +ᵥ 0` is an additive monoid homomorphism from `M` to `N`."] -def MonoidHom.smulOneHom {M N} [Monoid M] [MulOneClass N] [MulAction M N] [IsScalarTower M N N] : - M →* N where - toFun x := x • (1 : N) - map_one' := one_smul _ _ - map_mul' x y := by rw [smul_one_mul, smul_smul] - -/-- A monoid homomorphism between two monoids M and N can be equivalently specified by a -multiplicative action of M on N that is compatible with the multiplication on N. -/ -@[to_additive -"A monoid homomorphism between two additive monoids M and N can be equivalently -specified by an additive action of M on N that is compatible with the addition on N."] -def monoidHomEquivMulActionIsScalarTower (M N) [Monoid M] [Monoid N] : - (M →* N) ≃ {_inst : MulAction M N // IsScalarTower M N N} where - toFun f := ⟨MulAction.compHom N f, SMul.comp.isScalarTower _⟩ - invFun := fun ⟨_, _⟩ ↦ MonoidHom.smulOneHom - left_inv f := MonoidHom.ext fun m ↦ mul_one (f m) - right_inv := fun ⟨_, _⟩ ↦ Subtype.ext <| MulAction.ext <| funext₂ <| smul_one_smul N - end CompatibleScalar - -variable (α) - -/-- The monoid of endomorphisms. - -Note that this is generalized by `CategoryTheory.End` to categories other than `Type u`. -/ -protected def Function.End := α → α - -instance : Monoid (Function.End α) where - one := id - mul := (· ∘ ·) - mul_assoc _ _ _ := rfl - mul_one _ := rfl - one_mul _ := rfl - npow n f := f^[n] - npow_succ _ _ := Function.iterate_succ _ _ - -instance : Inhabited (Function.End α) := ⟨1⟩ - -variable {α} - -/-- The tautological action by `Function.End α` on `α`. - -This is generalized to bundled endomorphisms by: -* `Equiv.Perm.applyMulAction` -* `AddMonoid.End.applyDistribMulAction` -* `AddMonoid.End.applyModule` -* `AddAut.applyDistribMulAction` -* `MulAut.applyMulDistribMulAction` -* `LinearEquiv.applyDistribMulAction` -* `LinearMap.applyModule` -* `RingHom.applyMulSemiringAction` -* `RingAut.applyMulSemiringAction` -* `AlgEquiv.applyMulSemiringAction` --/ -instance Function.End.applyMulAction : MulAction (Function.End α) α where - smul := (· <| ·) - one_smul _ := rfl - mul_smul _ _ _ := rfl - -@[simp] lemma Function.End.smul_def (f : Function.End α) (a : α) : f • a = f a := rfl - ---TODO - This statement should be somethting like `toFun (f * g) = toFun f ∘ toFun g` -lemma Function.End.mul_def (f g : Function.End α) : (f * g) = f ∘ g := rfl - ---TODO - This statement should be somethting like `toFun 1 = id` -lemma Function.End.one_def : (1 : Function.End α) = id := rfl - -/-- `Function.End.applyMulAction` is faithful. -/ -instance Function.End.apply_FaithfulSMul : FaithfulSMul (Function.End α) α := - ⟨fun {_ _} ↦ funext⟩ - -/-- The monoid hom representing a monoid action. - -When `M` is a group, see `MulAction.toPermHom`. -/ -def MulAction.toEndHom [Monoid M] [MulAction M α] : M →* Function.End α where - toFun := (· • ·) - map_one' := funext (one_smul M) - map_mul' x y := funext (mul_smul x y) - -/-- The monoid action induced by a monoid hom to `Function.End α` - -See note [reducible non-instances]. -/ -abbrev MulAction.ofEndHom [Monoid M] (f : M →* Function.End α) : MulAction M α := - MulAction.compHom α f - -/-! ### `Additive`, `Multiplicative` -/ - -section - -open Additive Multiplicative - -instance Additive.vadd [SMul α β] : VAdd (Additive α) β where vadd a := (toMul a • ·) - -instance Multiplicative.smul [VAdd α β] : SMul (Multiplicative α) β where smul a := (toAdd a +ᵥ ·) - -@[simp] lemma toMul_smul [SMul α β] (a) (b : β) : (toMul a : α) • b = a +ᵥ b := rfl - -@[simp] lemma ofMul_vadd [SMul α β] (a : α) (b : β) : ofMul a +ᵥ b = a • b := rfl - -@[simp] lemma toAdd_vadd [VAdd α β] (a) (b : β) : (toAdd a : α) +ᵥ b = a • b := rfl - -@[simp] lemma ofAdd_smul [VAdd α β] (a : α) (b : β) : ofAdd a • b = a +ᵥ b := rfl - --- Porting note: I don't know why `one_smul` can do without an explicit α and `mul_smul` can't. -instance Additive.addAction [Monoid α] [MulAction α β] : AddAction (Additive α) β where - zero_vadd := MulAction.one_smul - add_vadd := MulAction.mul_smul (α := α) - -instance Multiplicative.mulAction [AddMonoid α] [AddAction α β] : - MulAction (Multiplicative α) β where - one_smul := AddAction.zero_vadd - mul_smul := AddAction.add_vadd (G := α) - -instance Additive.addAction_isPretransitive [Monoid α] [MulAction α β] - [MulAction.IsPretransitive α β] : AddAction.IsPretransitive (Additive α) β := - ⟨@MulAction.exists_smul_eq α _ _ _⟩ - -instance Multiplicative.mulAction_isPretransitive [AddMonoid α] [AddAction α β] - [AddAction.IsPretransitive α β] : MulAction.IsPretransitive (Multiplicative α) β := - ⟨@AddAction.exists_vadd_eq α _ _ _⟩ - -instance Additive.vaddCommClass [SMul α γ] [SMul β γ] [SMulCommClass α β γ] : - VAddCommClass (Additive α) (Additive β) γ := - ⟨@smul_comm α β _ _ _ _⟩ - -instance Multiplicative.smulCommClass [VAdd α γ] [VAdd β γ] [VAddCommClass α β γ] : - SMulCommClass (Multiplicative α) (Multiplicative β) γ := - ⟨@vadd_comm α β _ _ _ _⟩ - -end - -/-- The tautological additive action by `Additive (Function.End α)` on `α`. -/ -instance AddAction.functionEnd : AddAction (Additive (Function.End α)) α := inferInstance - -/-- The additive monoid hom representing an additive monoid action. - -When `M` is a group, see `AddAction.toPermHom`. -/ -def AddAction.toEndHom [AddMonoid M] [AddAction M α] : M →+ Additive (Function.End α) := - MonoidHom.toAdditive'' MulAction.toEndHom - -/-- The additive action induced by a hom to `Additive (Function.End α)` - -See note [reducible non-instances]. -/ -abbrev AddAction.ofEndHom [AddMonoid M] (f : M →+ Additive (Function.End α)) : AddAction M α := - AddAction.compHom α f diff --git a/Mathlib/Algebra/Group/Action/End.lean b/Mathlib/Algebra/Group/Action/End.lean new file mode 100644 index 0000000000000..302230a0828c8 --- /dev/null +++ b/Mathlib/Algebra/Group/Action/End.lean @@ -0,0 +1,169 @@ +/- +Copyright (c) 2018 Chris Hughes. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Chris Hughes, Yury Kudryashov +-/ +import Mathlib.Algebra.Group.Action.Defs +import Mathlib.Algebra.Group.Hom.Defs + +/-! +# Endomorphisms, homomorphisms and group actions + +This file defines `Function.End` as the monoid of endomorphisms on a type, +and provides results relating group actions with these endomorphisms. + +## Notation + +- `a • b` is used as notation for `SMul.smul a b`. +- `a +ᵥ b` is used as notation for `VAdd.vadd a b`. + +## Implementation details + +This file should avoid depending on other parts of `GroupTheory`, to avoid import cycles. +More sophisticated lemmas belong in `GroupTheory.GroupAction`. + +## Tags + +group action +-/ + +assert_not_exists MonoidWithZero + +open Function (Injective Surjective) + +variable {M N G H α β γ δ : Type*} + +section +variable [Monoid M] [MulAction M α] + +/-- Push forward the action of `R` on `M` along a compatible surjective map `f : R →* S`. + +See also `Function.Surjective.distribMulActionLeft` and `Function.Surjective.moduleLeft`. +-/ +@[to_additive +"Push forward the action of `R` on `M` along a compatible surjective map `f : R →+ S`."] +abbrev Function.Surjective.mulActionLeft {R S M : Type*} [Monoid R] [MulAction R M] [Monoid S] + [SMul S M] (f : R →* S) (hf : Surjective f) (hsmul : ∀ (c) (x : M), f c • x = c • x) : + MulAction S M where + smul := (· • ·) + one_smul b := by rw [← f.map_one, hsmul, one_smul] + mul_smul := hf.forall₂.mpr fun a b x ↦ by simp only [← f.map_mul, hsmul, mul_smul] + +namespace MulAction + +variable (α) + +/-- A multiplicative action of `M` on `α` and a monoid homomorphism `N → M` induce +a multiplicative action of `N` on `α`. + +See note [reducible non-instances]. -/ +@[to_additive] +abbrev compHom [Monoid N] (g : N →* M) : MulAction N α where + smul := SMul.comp.smul g + -- Porting note: was `by simp [g.map_one, MulAction.one_smul]` + one_smul _ := by simpa [(· • ·)] using MulAction.one_smul .. + -- Porting note: was `by simp [g.map_mul, MulAction.mul_smul]` + mul_smul _ _ _ := by simpa [(· • ·)] using MulAction.mul_smul .. + +/-- An additive action of `M` on `α` and an additive monoid homomorphism `N → M` induce +an additive action of `N` on `α`. + +See note [reducible non-instances]. -/ +add_decl_doc AddAction.compHom + +@[to_additive] +lemma compHom_smul_def + {E F G : Type*} [Monoid E] [Monoid F] [MulAction F G] (f : E →* F) (a : E) (x : G) : + letI : MulAction E G := MulAction.compHom _ f + a • x = (f a) • x := rfl + +end MulAction +end + +section CompatibleScalar + +/-- If the multiplicative action of `M` on `N` is compatible with multiplication on `N`, then +`fun x ↦ x • 1` is a monoid homomorphism from `M` to `N`. -/ +@[to_additive (attr := simps) +"If the additive action of `M` on `N` is compatible with addition on `N`, then +`fun x ↦ x +ᵥ 0` is an additive monoid homomorphism from `M` to `N`."] +def MonoidHom.smulOneHom {M N} [Monoid M] [MulOneClass N] [MulAction M N] [IsScalarTower M N N] : + M →* N where + toFun x := x • (1 : N) + map_one' := one_smul _ _ + map_mul' x y := by rw [smul_one_mul, smul_smul] + +/-- A monoid homomorphism between two monoids M and N can be equivalently specified by a +multiplicative action of M on N that is compatible with the multiplication on N. -/ +@[to_additive +"A monoid homomorphism between two additive monoids M and N can be equivalently +specified by an additive action of M on N that is compatible with the addition on N."] +def monoidHomEquivMulActionIsScalarTower (M N) [Monoid M] [Monoid N] : + (M →* N) ≃ {_inst : MulAction M N // IsScalarTower M N N} where + toFun f := ⟨MulAction.compHom N f, SMul.comp.isScalarTower _⟩ + invFun := fun ⟨_, _⟩ ↦ MonoidHom.smulOneHom + left_inv f := MonoidHom.ext fun m ↦ mul_one (f m) + right_inv := fun ⟨_, _⟩ ↦ Subtype.ext <| MulAction.ext <| funext₂ <| smul_one_smul N + +end CompatibleScalar + +variable (α) + +/-- The monoid of endomorphisms. + +Note that this is generalized by `CategoryTheory.End` to categories other than `Type u`. -/ +protected def Function.End := α → α + +instance : Monoid (Function.End α) where + one := id + mul := (· ∘ ·) + mul_assoc _ _ _ := rfl + mul_one _ := rfl + one_mul _ := rfl + npow n f := f^[n] + npow_succ _ _ := Function.iterate_succ _ _ + +instance : Inhabited (Function.End α) := ⟨1⟩ + +variable {α} + +/-- The tautological action by `Function.End α` on `α`. + +This is generalized to bundled endomorphisms by: +* `Equiv.Perm.applyMulAction` +* `AddMonoid.End.applyDistribMulAction` +* `AddMonoid.End.applyModule` +* `AddAut.applyDistribMulAction` +* `MulAut.applyMulDistribMulAction` +* `LinearEquiv.applyDistribMulAction` +* `LinearMap.applyModule` +* `RingHom.applyMulSemiringAction` +* `RingAut.applyMulSemiringAction` +* `AlgEquiv.applyMulSemiringAction` +-/ +instance Function.End.applyMulAction : MulAction (Function.End α) α where + smul := (· <| ·) + one_smul _ := rfl + mul_smul _ _ _ := rfl + +@[simp] lemma Function.End.smul_def (f : Function.End α) (a : α) : f • a = f a := rfl + +--TODO - This statement should be somethting like `toFun (f * g) = toFun f ∘ toFun g` +lemma Function.End.mul_def (f g : Function.End α) : (f * g) = f ∘ g := rfl + +--TODO - This statement should be somethting like `toFun 1 = id` +lemma Function.End.one_def : (1 : Function.End α) = id := rfl + +/-- The monoid hom representing a monoid action. + +When `M` is a group, see `MulAction.toPermHom`. -/ +def MulAction.toEndHom [Monoid M] [MulAction M α] : M →* Function.End α where + toFun := (· • ·) + map_one' := funext (one_smul M) + map_mul' x y := funext (mul_smul x y) + +/-- The monoid action induced by a monoid hom to `Function.End α` + +See note [reducible non-instances]. -/ +abbrev MulAction.ofEndHom [Monoid M] (f : M →* Function.End α) : MulAction M α := + MulAction.compHom α f diff --git a/Mathlib/Algebra/Group/Action/Faithful.lean b/Mathlib/Algebra/Group/Action/Faithful.lean new file mode 100644 index 0000000000000..19ba00c7bcc60 --- /dev/null +++ b/Mathlib/Algebra/Group/Action/Faithful.lean @@ -0,0 +1,61 @@ +/- +Copyright (c) 2018 Chris Hughes. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Chris Hughes, Yury Kudryashov +-/ +import Mathlib.Algebra.Group.Action.End + +/-! +# Faithful group actions + +This file provides typeclasses for faithful actions. + +## Notation + +- `a • b` is used as notation for `SMul.smul a b`. +- `a +ᵥ b` is used as notation for `VAdd.vadd a b`. + +## Implementation details + +This file should avoid depending on other parts of `GroupTheory`, to avoid import cycles. +More sophisticated lemmas belong in `GroupTheory.GroupAction`. + +## Tags + +group action +-/ + +assert_not_exists MonoidWithZero + +open Function (Injective Surjective) + +variable {M N G H α β γ δ : Type*} + +/-! ### Faithful actions -/ + +/-- Typeclass for faithful actions. -/ +class FaithfulVAdd (G : Type*) (P : Type*) [VAdd G P] : Prop where + /-- Two elements `g₁` and `g₂` are equal whenever they act in the same way on all points. -/ + eq_of_vadd_eq_vadd : ∀ {g₁ g₂ : G}, (∀ p : P, g₁ +ᵥ p = g₂ +ᵥ p) → g₁ = g₂ + +/-- Typeclass for faithful actions. -/ +@[to_additive] +class FaithfulSMul (M : Type*) (α : Type*) [SMul M α] : Prop where + /-- Two elements `m₁` and `m₂` are equal whenever they act in the same way on all points. -/ + eq_of_smul_eq_smul : ∀ {m₁ m₂ : M}, (∀ a : α, m₁ • a = m₂ • a) → m₁ = m₂ + +export FaithfulSMul (eq_of_smul_eq_smul) +export FaithfulVAdd (eq_of_vadd_eq_vadd) + +@[to_additive] +lemma smul_left_injective' [SMul M α] [FaithfulSMul M α] : Injective ((· • ·) : M → α → α) := + fun _ _ h ↦ FaithfulSMul.eq_of_smul_eq_smul (congr_fun h) + +/-- `Monoid.toMulAction` is faithful on cancellative monoids. -/ +@[to_additive " `AddMonoid.toAddAction` is faithful on additive cancellative monoids. "] +instance RightCancelMonoid.faithfulSMul [RightCancelMonoid α] : FaithfulSMul α α := + ⟨fun h ↦ mul_right_cancel (h 1)⟩ + +/-- `Function.End.applyMulAction` is faithful. -/ +instance Function.End.apply_FaithfulSMul : FaithfulSMul (Function.End α) α := + ⟨fun {_ _} ↦ funext⟩ diff --git a/Mathlib/Algebra/Group/Action/Opposite.lean b/Mathlib/Algebra/Group/Action/Opposite.lean index 1217487ac1eda..1b5ea20d99733 100644 --- a/Mathlib/Algebra/Group/Action/Opposite.lean +++ b/Mathlib/Algebra/Group/Action/Opposite.lean @@ -3,7 +3,8 @@ Copyright (c) 2020 Eric Wieser. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Eric Wieser -/ -import Mathlib.Algebra.Group.Action.Defs +import Mathlib.Algebra.Group.Action.Faithful +import Mathlib.Algebra.Group.Action.Pretransitive import Mathlib.Algebra.Group.Opposite /-! diff --git a/Mathlib/Algebra/Group/Action/Option.lean b/Mathlib/Algebra/Group/Action/Option.lean index 1a56618022160..72a1fef15f92a 100644 --- a/Mathlib/Algebra/Group/Action/Option.lean +++ b/Mathlib/Algebra/Group/Action/Option.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ -import Mathlib.Algebra.Group.Action.Defs +import Mathlib.Algebra.Group.Action.Faithful /-! # Option instances for additive and multiplicative actions diff --git a/Mathlib/Algebra/Group/Action/Pi.lean b/Mathlib/Algebra/Group/Action/Pi.lean index f1ed601c36220..7e61da3072d3a 100644 --- a/Mathlib/Algebra/Group/Action/Pi.lean +++ b/Mathlib/Algebra/Group/Action/Pi.lean @@ -3,7 +3,7 @@ Copyright (c) 2018 Simon Hudon. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Simon Hudon, Patrick Massot -/ -import Mathlib.Algebra.Group.Action.Defs +import Mathlib.Algebra.Group.Action.Faithful import Mathlib.Data.Set.Function /-! diff --git a/Mathlib/Algebra/Group/Action/Pretransitive.lean b/Mathlib/Algebra/Group/Action/Pretransitive.lean new file mode 100644 index 0000000000000..4741b1857a579 --- /dev/null +++ b/Mathlib/Algebra/Group/Action/Pretransitive.lean @@ -0,0 +1,130 @@ +/- +Copyright (c) 2018 Chris Hughes. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Chris Hughes, Yury Kudryashov +-/ +import Mathlib.Algebra.Group.Action.TypeTags + +/-! +# Pretransitive group actions + +This file defines a typeclass for pretransitive group actions. + +## Notation + +- `a • b` is used as notation for `SMul.smul a b`. +- `a +ᵥ b` is used as notation for `VAdd.vadd a b`. + +## Implementation details + +This file should avoid depending on other parts of `GroupTheory`, to avoid import cycles. +More sophisticated lemmas belong in `GroupTheory.GroupAction`. + +## Tags + +group action +-/ + +assert_not_exists MonoidWithZero + +open Function (Injective Surjective) + +variable {M N G H α β γ δ : Type*} + +/-! +### (Pre)transitive action + +`M` acts pretransitively on `α` if for any `x y` there is `g` such that `g • x = y` (or `g +ᵥ x = y` +for an additive action). A transitive action should furthermore have `α` nonempty. + +In this section we define typeclasses `MulAction.IsPretransitive` and +`AddAction.IsPretransitive` and provide `MulAction.exists_smul_eq`/`AddAction.exists_vadd_eq`, +`MulAction.surjective_smul`/`AddAction.surjective_vadd` as public interface to access this +property. We do not provide typeclasses `*Action.IsTransitive`; users should assume +`[MulAction.IsPretransitive M α] [Nonempty α]` instead. +-/ + +/-- `M` acts pretransitively on `α` if for any `x y` there is `g` such that `g +ᵥ x = y`. + A transitive action should furthermore have `α` nonempty. -/ +class AddAction.IsPretransitive (M α : Type*) [VAdd M α] : Prop where + /-- There is `g` such that `g +ᵥ x = y`. -/ + exists_vadd_eq : ∀ x y : α, ∃ g : M, g +ᵥ x = y + +/-- `M` acts pretransitively on `α` if for any `x y` there is `g` such that `g • x = y`. + A transitive action should furthermore have `α` nonempty. -/ +@[to_additive] +class MulAction.IsPretransitive (M α : Type*) [SMul M α] : Prop where + /-- There is `g` such that `g • x = y`. -/ + exists_smul_eq : ∀ x y : α, ∃ g : M, g • x = y + +namespace MulAction +variable (M) [SMul M α] [IsPretransitive M α] + +@[to_additive] +lemma exists_smul_eq (x y : α) : ∃ m : M, m • x = y := IsPretransitive.exists_smul_eq x y + +@[to_additive] +lemma surjective_smul (x : α) : Surjective fun c : M ↦ c • x := exists_smul_eq M x + +/-- The regular action of a group on itself is transitive. -/ +@[to_additive "The regular action of a group on itself is transitive."] +instance Regular.isPretransitive [Group G] : IsPretransitive G G := + ⟨fun x y ↦ ⟨y * x⁻¹, inv_mul_cancel_right _ _⟩⟩ + +end MulAction + +namespace MulAction + +variable [Monoid M] [MulAction M α] + +variable (α) + +/-- If an action is transitive, then composing this action with a surjective homomorphism gives +again a transitive action. -/ +@[to_additive] +lemma isPretransitive_compHom {E F G : Type*} [Monoid E] [Monoid F] [MulAction F G] + [IsPretransitive F G] {f : E →* F} (hf : Surjective f) : + letI : MulAction E G := MulAction.compHom _ f + IsPretransitive E G := by + let _ : MulAction E G := MulAction.compHom _ f + refine ⟨fun x y ↦ ?_⟩ + obtain ⟨m, rfl⟩ : ∃ m : F, m • x = y := exists_smul_eq F x y + obtain ⟨e, rfl⟩ : ∃ e, f e = m := hf m + exact ⟨e, rfl⟩ + +@[to_additive] +lemma IsPretransitive.of_smul_eq {M N α : Type*} [SMul M α] [SMul N α] [IsPretransitive M α] + (f : M → N) (hf : ∀ {c : M} {x : α}, f c • x = c • x) : IsPretransitive N α where + exists_smul_eq x y := (exists_smul_eq x y).elim fun m h ↦ ⟨f m, hf.trans h⟩ + +@[to_additive] +lemma IsPretransitive.of_compHom {M N α : Type*} [Monoid M] [Monoid N] [MulAction N α] + (f : M →* N) [h : letI := compHom α f; IsPretransitive M α] : IsPretransitive N α := + letI := compHom α f; h.of_smul_eq f rfl + +end MulAction + +section CompatibleScalar + +@[to_additive] +lemma MulAction.IsPretransitive.of_isScalarTower (M : Type*) {N α : Type*} [Monoid N] [SMul M N] + [MulAction N α] [SMul M α] [IsScalarTower M N α] [IsPretransitive M α] : IsPretransitive N α := + of_smul_eq (fun x : M ↦ x • 1) (smul_one_smul N _ _) + +end CompatibleScalar + +/-! ### `Additive`, `Multiplicative` -/ + +section + +open Additive Multiplicative + +instance Additive.addAction_isPretransitive [Monoid α] [MulAction α β] + [MulAction.IsPretransitive α β] : AddAction.IsPretransitive (Additive α) β := + ⟨@MulAction.exists_smul_eq α _ _ _⟩ + +instance Multiplicative.mulAction_isPretransitive [AddMonoid α] [AddAction α β] + [AddAction.IsPretransitive α β] : MulAction.IsPretransitive (Multiplicative α) β := + ⟨@AddAction.exists_vadd_eq α _ _ _⟩ + +end diff --git a/Mathlib/Algebra/Group/Action/Prod.lean b/Mathlib/Algebra/Group/Action/Prod.lean index e4434083c689b..898caa223cdc3 100644 --- a/Mathlib/Algebra/Group/Action/Prod.lean +++ b/Mathlib/Algebra/Group/Action/Prod.lean @@ -3,7 +3,7 @@ Copyright (c) 2018 Simon Hudon. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Simon Hudon, Patrick Massot, Eric Wieser -/ -import Mathlib.Algebra.Group.Action.Defs +import Mathlib.Algebra.Group.Action.Faithful import Mathlib.Algebra.Group.Prod /-! diff --git a/Mathlib/Algebra/Group/Action/Sigma.lean b/Mathlib/Algebra/Group/Action/Sigma.lean index 5aaeee6501dcc..bdd561079350c 100644 --- a/Mathlib/Algebra/Group/Action/Sigma.lean +++ b/Mathlib/Algebra/Group/Action/Sigma.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ -import Mathlib.Algebra.Group.Action.Defs +import Mathlib.Algebra.Group.Action.Faithful /-! # Sigma instances for additive and multiplicative actions diff --git a/Mathlib/Algebra/Group/Action/Sum.lean b/Mathlib/Algebra/Group/Action/Sum.lean index 0286b8e19b6e9..d7db12fe9d54b 100644 --- a/Mathlib/Algebra/Group/Action/Sum.lean +++ b/Mathlib/Algebra/Group/Action/Sum.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ -import Mathlib.Algebra.Group.Action.Defs +import Mathlib.Algebra.Group.Action.Faithful /-! # Sum instances for additive and multiplicative actions diff --git a/Mathlib/Algebra/Group/Action/TypeTags.lean b/Mathlib/Algebra/Group/Action/TypeTags.lean new file mode 100644 index 0000000000000..23a3579798c6e --- /dev/null +++ b/Mathlib/Algebra/Group/Action/TypeTags.lean @@ -0,0 +1,72 @@ +/- +Copyright (c) 2018 Chris Hughes. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Chris Hughes, Yury Kudryashov +-/ +import Mathlib.Algebra.Group.Action.End +import Mathlib.Algebra.Group.TypeTags + +/-! +# Additive and Multiplicative for group actions + +## Tags + +group action +-/ + +assert_not_exists MonoidWithZero + +open Function (Injective Surjective) + +variable {M N G H α β γ δ : Type*} + +section + +open Additive Multiplicative + +instance Additive.vadd [SMul α β] : VAdd (Additive α) β where vadd a := (toMul a • ·) + +instance Multiplicative.smul [VAdd α β] : SMul (Multiplicative α) β where smul a := (toAdd a +ᵥ ·) + +@[simp] lemma toMul_smul [SMul α β] (a) (b : β) : (toMul a : α) • b = a +ᵥ b := rfl + +@[simp] lemma ofMul_vadd [SMul α β] (a : α) (b : β) : ofMul a +ᵥ b = a • b := rfl + +@[simp] lemma toAdd_vadd [VAdd α β] (a) (b : β) : (toAdd a : α) +ᵥ b = a • b := rfl + +@[simp] lemma ofAdd_smul [VAdd α β] (a : α) (b : β) : ofAdd a • b = a +ᵥ b := rfl + +-- Porting note: I don't know why `one_smul` can do without an explicit α and `mul_smul` can't. +instance Additive.addAction [Monoid α] [MulAction α β] : AddAction (Additive α) β where + zero_vadd := MulAction.one_smul + add_vadd := MulAction.mul_smul (α := α) + +instance Multiplicative.mulAction [AddMonoid α] [AddAction α β] : + MulAction (Multiplicative α) β where + one_smul := AddAction.zero_vadd + mul_smul := AddAction.add_vadd (G := α) + +instance Additive.vaddCommClass [SMul α γ] [SMul β γ] [SMulCommClass α β γ] : + VAddCommClass (Additive α) (Additive β) γ := + ⟨@smul_comm α β _ _ _ _⟩ + +instance Multiplicative.smulCommClass [VAdd α γ] [VAdd β γ] [VAddCommClass α β γ] : + SMulCommClass (Multiplicative α) (Multiplicative β) γ := + ⟨@vadd_comm α β _ _ _ _⟩ + +end + +/-- The tautological additive action by `Additive (Function.End α)` on `α`. -/ +instance AddAction.functionEnd : AddAction (Additive (Function.End α)) α := inferInstance + +/-- The additive monoid hom representing an additive monoid action. + +When `M` is a group, see `AddAction.toPermHom`. -/ +def AddAction.toEndHom [AddMonoid M] [AddAction M α] : M →+ Additive (Function.End α) := + MonoidHom.toAdditive'' MulAction.toEndHom + +/-- The additive action induced by a hom to `Additive (Function.End α)` + +See note [reducible non-instances]. -/ +abbrev AddAction.ofEndHom [AddMonoid M] (f : M →+ Additive (Function.End α)) : AddAction M α := + AddAction.compHom α f diff --git a/Mathlib/Algebra/Group/Action/Units.lean b/Mathlib/Algebra/Group/Action/Units.lean index 9092b7ac468d8..f2a6ef61171a5 100644 --- a/Mathlib/Algebra/Group/Action/Units.lean +++ b/Mathlib/Algebra/Group/Action/Units.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Eric Wieser. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Eric Wieser -/ -import Mathlib.Algebra.Group.Action.Defs +import Mathlib.Algebra.Group.Action.Faithful import Mathlib.Algebra.Group.Basic import Mathlib.Algebra.Group.Units.Defs diff --git a/Mathlib/Algebra/Group/Submonoid/Operations.lean b/Mathlib/Algebra/Group/Submonoid/Operations.lean index 300498e17328b..e5ca00f78346e 100644 --- a/Mathlib/Algebra/Group/Submonoid/Operations.lean +++ b/Mathlib/Algebra/Group/Submonoid/Operations.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Kenny Lau, Johan Commelin, Mario Carneiro, Kevin Buzzard, Amelia Livingston, Yury Kudryashov -/ -import Mathlib.Algebra.Group.Action.Defs +import Mathlib.Algebra.Group.Action.Faithful import Mathlib.Algebra.Group.Nat import Mathlib.Algebra.Group.Submonoid.Basic import Mathlib.Algebra.Group.Subsemigroup.Operations From a3c1388b7856689e7fffbe49cdffb1395c61a54d Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Thu, 24 Oct 2024 20:59:22 +0000 Subject: [PATCH 394/505] chore(Algebra/CharP): reduce theory requirements for `CharP.Defs` (#18182) The `CharP/Defs.lean` file imports some structures and proves quite a bit of theory. This PR tries to ensure that the focus is on the definitions instead, namely `CharP`, `ringChar`, `ExpChar` and `ringExpChar`. In the process, I made the following moves: * `CharP/Basic.lean`: renamed to `CharP/Lemmas.lean` * `CharP/Defs.lean`: all the lemmas that are not immediate consequences, or not needed to define `ringChar` or `expChar` have been moved to the new file `CharP/Basic.lean` Note that `CharP/Basic.lean` and `CharP/Lemmas.lean` can be imported independently. They are both unstructured collections of results so I'm not sure how we want to clean that up further. --- Counterexamples/CharPZeroNeCharZero.lean | 2 +- Mathlib.lean | 1 + Mathlib/Algebra/CharP/Algebra.lean | 2 +- Mathlib/Algebra/CharP/Basic.lean | 470 +++++------------- Mathlib/Algebra/CharP/CharAndCard.lean | 1 + Mathlib/Algebra/CharP/Defs.lean | 140 +----- Mathlib/Algebra/CharP/ExpChar.lean | 2 +- Mathlib/Algebra/CharP/Lemmas.lean | 391 +++++++++++++++ Mathlib/Algebra/CharP/Reduced.lean | 2 +- Mathlib/Algebra/CharP/Two.lean | 2 +- Mathlib/Algebra/Polynomial/Expand.lean | 2 +- .../Combinatorics/Additive/FreimanHom.lean | 2 +- Mathlib/Data/Nat/Choose/Lucas.lean | 2 +- Mathlib/Data/ZMod/Basic.lean | 2 +- Mathlib/Data/ZMod/Defs.lean | 4 +- Mathlib/ModelTheory/Algebra/Field/CharP.lean | 3 +- Mathlib/NumberTheory/MulChar/Basic.lean | 1 + Mathlib/RingTheory/Polynomial/Dickson.lean | 2 +- 18 files changed, 548 insertions(+), 483 deletions(-) create mode 100644 Mathlib/Algebra/CharP/Lemmas.lean diff --git a/Counterexamples/CharPZeroNeCharZero.lean b/Counterexamples/CharPZeroNeCharZero.lean index 06e30c54a4a82..4edc1ed83998c 100644 --- a/Counterexamples/CharPZeroNeCharZero.lean +++ b/Counterexamples/CharPZeroNeCharZero.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Damiano Testa. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Damiano Testa, Eric Wieser -/ -import Mathlib.Algebra.CharP.Basic +import Mathlib.Algebra.CharP.Lemmas import Mathlib.Algebra.PUnitInstances.Algebra /-! # `CharP R 0` and `CharZero R` need not coincide for semirings diff --git a/Mathlib.lean b/Mathlib.lean index 22dc00ac223e7..6d9a2bf44e45d 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -157,6 +157,7 @@ import Mathlib.Algebra.CharP.Defs import Mathlib.Algebra.CharP.ExpChar import Mathlib.Algebra.CharP.IntermediateField import Mathlib.Algebra.CharP.Invertible +import Mathlib.Algebra.CharP.Lemmas import Mathlib.Algebra.CharP.LocalRing import Mathlib.Algebra.CharP.MixedCharZero import Mathlib.Algebra.CharP.Pi diff --git a/Mathlib/Algebra/CharP/Algebra.lean b/Mathlib/Algebra/CharP/Algebra.lean index 7d73460eb9f55..4b23ad5f49507 100644 --- a/Mathlib/Algebra/CharP/Algebra.lean +++ b/Mathlib/Algebra/CharP/Algebra.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Jon Eugster. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jon Eugster, Eric Wieser -/ -import Mathlib.Algebra.CharP.Defs +import Mathlib.Algebra.CharP.Basic import Mathlib.Algebra.FreeAlgebra import Mathlib.RingTheory.Localization.FractionRing diff --git a/Mathlib/Algebra/CharP/Basic.lean b/Mathlib/Algebra/CharP/Basic.lean index 271aa60a6d144..bb0db6f85c58a 100644 --- a/Mathlib/Algebra/CharP/Basic.lean +++ b/Mathlib/Algebra/CharP/Basic.lean @@ -4,388 +4,184 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau, Joey van Langen, Casper Putz -/ import Mathlib.Algebra.CharP.Defs -import Mathlib.Algebra.GroupPower.IterateHom -import Mathlib.Data.Nat.Multiplicity -import Mathlib.Data.Nat.Choose.Sum +import Mathlib.Algebra.Field.Basic +import Mathlib.Algebra.Group.Fin.Basic +import Mathlib.Algebra.Group.ULift +import Mathlib.Data.Int.ModEq +import Mathlib.Data.Nat.Cast.Prod +import Mathlib.Data.Nat.Find +import Mathlib.Data.Nat.Prime.Defs +import Mathlib.Data.ULift +import Mathlib.Tactic.NormNum.Basic +import Mathlib.Order.Interval.Set.Basic /-! # Characteristic of semirings --/ - -assert_not_exists Algebra -assert_not_exists LinearMap -assert_not_exists orderOf - -open Finset - -variable {R S : Type*} - -namespace Commute - -variable [Semiring R] {p : ℕ} (hp : p.Prime) {x y : R} -include hp - -protected theorem add_pow_prime_pow_eq (h : Commute x y) (n : ℕ) : - (x + y) ^ p ^ n = - x ^ p ^ n + y ^ p ^ n + - p * ∑ k ∈ Ioo 0 (p ^ n), x ^ k * y ^ (p ^ n - k) * ↑((p ^ n).choose k / p) := by - trans x ^ p ^ n + y ^ p ^ n + ∑ k ∈ Ioo 0 (p ^ n), x ^ k * y ^ (p ^ n - k) * (p ^ n).choose k - · simp_rw [h.add_pow, ← Nat.Ico_zero_eq_range, Nat.Ico_succ_right, Icc_eq_cons_Ico (zero_le _), - Finset.sum_cons, Ico_eq_cons_Ioo (pow_pos hp.pos _), Finset.sum_cons, tsub_self, tsub_zero, - pow_zero, Nat.choose_zero_right, Nat.choose_self, Nat.cast_one, mul_one, one_mul, ← add_assoc] - · congr 1 - simp_rw [Finset.mul_sum, Nat.cast_comm, mul_assoc _ _ (p : R), ← Nat.cast_mul] - refine Finset.sum_congr rfl fun i hi => ?_ - rw [mem_Ioo] at hi - rw [Nat.div_mul_cancel (hp.dvd_choose_pow hi.1.ne' hi.2.ne)] - -protected theorem add_pow_prime_eq (h : Commute x y) : - (x + y) ^ p = - x ^ p + y ^ p + p * ∑ k ∈ Finset.Ioo 0 p, x ^ k * y ^ (p - k) * ↑(p.choose k / p) := by - simpa using h.add_pow_prime_pow_eq hp 1 - -protected theorem exists_add_pow_prime_pow_eq (h : Commute x y) (n : ℕ) : - ∃ r, (x + y) ^ p ^ n = x ^ p ^ n + y ^ p ^ n + p * r := - ⟨_, h.add_pow_prime_pow_eq hp n⟩ - -protected theorem exists_add_pow_prime_eq (h : Commute x y) : - ∃ r, (x + y) ^ p = x ^ p + y ^ p + p * r := - ⟨_, h.add_pow_prime_eq hp⟩ - -end Commute - -section CommSemiring - -variable [CommSemiring R] {p : ℕ} (hp : p.Prime) (x y : R) (n : ℕ) -include hp - -theorem add_pow_prime_pow_eq : - (x + y) ^ p ^ n = - x ^ p ^ n + y ^ p ^ n + - p * ∑ k ∈ Finset.Ioo 0 (p ^ n), x ^ k * y ^ (p ^ n - k) * ↑((p ^ n).choose k / p) := - (Commute.all x y).add_pow_prime_pow_eq hp n - -theorem add_pow_prime_eq : - (x + y) ^ p = - x ^ p + y ^ p + p * ∑ k ∈ Finset.Ioo 0 p, x ^ k * y ^ (p - k) * ↑(p.choose k / p) := - (Commute.all x y).add_pow_prime_eq hp - -theorem exists_add_pow_prime_pow_eq : - ∃ r, (x + y) ^ p ^ n = x ^ p ^ n + y ^ p ^ n + p * r := - (Commute.all x y).exists_add_pow_prime_pow_eq hp n - -theorem exists_add_pow_prime_eq : - ∃ r, (x + y) ^ p = x ^ p + y ^ p + p * r := - (Commute.all x y).exists_add_pow_prime_eq hp - -end CommSemiring - -section Semiring -variable [Semiring R] {x y : R} (p n : ℕ) - -section ExpChar -variable [hR : ExpChar R p] - -lemma add_pow_expChar_of_commute (h : Commute x y) : (x + y) ^ p = x ^ p + y ^ p := by - obtain _ | hprime := hR - · simp only [pow_one] - · let ⟨r, hr⟩ := h.exists_add_pow_prime_eq hprime - simp [hr] - -lemma add_pow_expChar_pow_of_commute (h : Commute x y) : - (x + y) ^ p ^ n = x ^ p ^ n + y ^ p ^ n := by - obtain _ | hprime := hR - · simp only [one_pow, pow_one] - · let ⟨r, hr⟩ := h.exists_add_pow_prime_pow_eq hprime n - simp [hr] - -lemma add_pow_eq_mul_pow_add_pow_div_expChar_of_commute (h : Commute x y) : - (x + y) ^ n = (x + y) ^ (n % p) * (x ^ p + y ^ p) ^ (n / p) := by - rw [← add_pow_expChar_of_commute _ h, ← pow_mul, ← pow_add, Nat.mod_add_div] - -end ExpChar - -section CharP -variable [hp : Fact p.Prime] [CharP R p] - -lemma add_pow_char_of_commute (h : Commute x y) : (x + y) ^ p = x ^ p + y ^ p := - add_pow_expChar_of_commute _ h - -lemma add_pow_char_pow_of_commute (h : Commute x y) : (x + y) ^ p ^ n = x ^ p ^ n + y ^ p ^ n := - add_pow_expChar_pow_of_commute _ _ h - -lemma add_pow_eq_mul_pow_add_pow_div_char_of_commute (h : Commute x y) : - (x + y) ^ n = (x + y) ^ (n % p) * (x ^ p + y ^ p) ^ (n / p) := - add_pow_eq_mul_pow_add_pow_div_expChar_of_commute _ _ h - -end CharP -end Semiring - -section CommSemiring -variable [CommSemiring R] (x y : R) (p n : ℕ) - -section ExpChar -variable [hR : ExpChar R p] - -lemma add_pow_expChar : (x + y) ^ p = x ^ p + y ^ p := add_pow_expChar_of_commute _ <| .all .. - -lemma add_pow_expChar_pow : (x + y) ^ p ^ n = x ^ p ^ n + y ^ p ^ n := - add_pow_expChar_pow_of_commute _ _ <| .all .. - -lemma add_pow_eq_mul_pow_add_pow_div_expChar : - (x + y) ^ n = (x + y) ^ (n % p) * (x ^ p + y ^ p) ^ (n / p) := - add_pow_eq_mul_pow_add_pow_div_expChar_of_commute _ _ <| .all .. - -end ExpChar - -section CharP -variable [hp : Fact p.Prime] [CharP R p] - -lemma add_pow_char : (x + y) ^ p = x ^ p + y ^ p := add_pow_expChar .. -lemma add_pow_char_pow : (x + y) ^ p ^ n = x ^ p ^ n + y ^ p ^ n := add_pow_expChar_pow .. - -lemma add_pow_eq_mul_pow_add_pow_div_char : - (x + y) ^ n = (x + y) ^ (n % p) * (x ^ p + y ^ p) ^ (n / p) := - add_pow_eq_mul_pow_add_pow_div_expChar .. - -@[deprecated (since := "2024-10-21")] -alias add_pow_eq_add_pow_mod_mul_pow_add_pow_div := add_pow_eq_mul_pow_add_pow_div_char - -end CharP -end CommSemiring - -section Ring -variable [Ring R] {x y : R} (p n : ℕ) - -section ExpChar -variable [hR : ExpChar R p] -include hR -lemma sub_pow_expChar_of_commute (h : Commute x y) : (x - y) ^ p = x ^ p - y ^ p := by - simp [eq_sub_iff_add_eq, ← add_pow_expChar_of_commute _ (h.sub_left rfl)] +This file collects some fundamental results on the characteristic of rings that don't need the extra +imports of `CharP/Lemmas.lean`. -lemma sub_pow_expChar_pow_of_commute (h : Commute x y) : - (x - y) ^ p ^ n = x ^ p ^ n - y ^ p ^ n := by - simp [eq_sub_iff_add_eq, ← add_pow_expChar_pow_of_commute _ _ (h.sub_left rfl)] - -lemma sub_pow_eq_mul_pow_sub_pow_div_expChar_of_commute (h : Commute x y) : - (x - y) ^ n = (x - y) ^ (n % p) * (x ^ p - y ^ p) ^ (n / p) := by - rw [← sub_pow_expChar_of_commute _ h, ← pow_mul, ← pow_add, Nat.mod_add_div] - -variable (R) +As such, we can probably reorganize and find a better home for most of these lemmas. +-/ -lemma neg_one_pow_expChar : (-1 : R) ^ p = -1 := by - rw [eq_neg_iff_add_eq_zero] - nth_rw 2 [← one_pow p] - rw [← add_pow_expChar_of_commute _ (Commute.one_right _), neg_add_cancel, - zero_pow (expChar_ne_zero R p)] +assert_not_exists Finset -lemma neg_one_pow_expChar_pow : (-1 : R) ^ p ^ n = -1 := by - rw [eq_neg_iff_add_eq_zero] - nth_rw 2 [← one_pow (p ^ n)] - rw [← add_pow_expChar_pow_of_commute _ _ (Commute.one_right _), neg_add_cancel, - zero_pow (pow_ne_zero _ <| expChar_ne_zero R p)] +open Set -end ExpChar +variable (R : Type*) -section CharP -variable [hp : Fact p.Prime] [CharP R p] +namespace CharP +section AddMonoidWithOne +variable [AddMonoidWithOne R] (p : ℕ) -lemma sub_pow_char_of_commute (h : Commute x y) : (x - y) ^ p = x ^ p - y ^ p := - sub_pow_expChar_of_commute _ h +variable [CharP R p] {a b : ℕ} -lemma sub_pow_char_pow_of_commute (h : Commute x y) : (x - y) ^ p ^ n = x ^ p ^ n - y ^ p ^ n := - sub_pow_expChar_pow_of_commute _ _ h +variable [IsRightCancelAdd R] -variable (R) +lemma natCast_injOn_Iio : (Set.Iio p).InjOn ((↑) : ℕ → R) := + fun _a ha _b hb hab ↦ ((natCast_eq_natCast _ _).1 hab).eq_of_lt_of_lt ha hb -lemma neg_one_pow_char : (-1 : R) ^ p = -1 := neg_one_pow_expChar .. +end AddMonoidWithOne -lemma neg_one_pow_char_pow : (-1 : R) ^ p ^ n = -1 := neg_one_pow_expChar_pow .. +section AddGroupWithOne +variable [AddGroupWithOne R] (p : ℕ) [CharP R p] {a b : ℤ} -lemma sub_pow_eq_mul_pow_sub_pow_div_char_of_commute (h : Commute x y) : - (x - y) ^ n = (x - y) ^ (n % p) * (x ^ p - y ^ p) ^ (n / p) := - sub_pow_eq_mul_pow_sub_pow_div_expChar_of_commute _ _ h +lemma intCast_injOn_Ico [IsRightCancelAdd R] : InjOn (Int.cast : ℤ → R) (Ico 0 p) := by + rintro a ⟨ha₀, ha⟩ b ⟨hb₀, hb⟩ hab + lift a to ℕ using ha₀ + lift b to ℕ using hb₀ + norm_cast at * + exact natCast_injOn_Iio _ _ ha hb hab +end AddGroupWithOne end CharP -end Ring - -section CommRing -variable [CommRing R] (x y : R) (n : ℕ) {p : ℕ} - -section ExpChar -variable [hR : ExpChar R p] -lemma sub_pow_expChar : (x - y) ^ p = x ^ p - y ^ p := sub_pow_expChar_of_commute _ <| .all .. +lemma RingHom.charP_iff_charP {K L : Type*} [DivisionRing K] [Semiring L] [Nontrivial L] + (f : K →+* L) (p : ℕ) : CharP K p ↔ CharP L p := by + simp only [charP_iff, ← f.injective.eq_iff, map_natCast f, map_zero f] -lemma sub_pow_expChar_pow : (x - y) ^ p ^ n = x ^ p ^ n - y ^ p ^ n := - sub_pow_expChar_pow_of_commute _ _ <| .all .. - -lemma sub_pow_eq_mul_pow_sub_pow_div_expChar : - (x - y) ^ n = (x - y) ^ (n % p) * (x ^ p - y ^ p) ^ (n / p) := - sub_pow_eq_mul_pow_sub_pow_div_expChar_of_commute _ _ <| .all .. - -end ExpChar - -section CharP -variable [hp : Fact p.Prime] [CharP R p] +namespace CharP -lemma sub_pow_char : (x - y) ^ p = x ^ p - y ^ p := sub_pow_expChar .. -lemma sub_pow_char_pow : (x - y) ^ p ^ n = x ^ p ^ n - y ^ p ^ n := sub_pow_expChar_pow .. +section NonAssocSemiring -lemma sub_pow_eq_mul_pow_sub_pow_div_char : - (x - y) ^ n = (x - y) ^ (n % p) * (x ^ p - y ^ p) ^ (n / p) := - sub_pow_eq_mul_pow_sub_pow_div_expChar .. +variable {R} [NonAssocSemiring R] +variable (R) in +/-- If a ring `R` is of characteristic `p`, then for any prime number `q` different from `p`, +it is not zero in `R`. -/ +lemma cast_ne_zero_of_ne_of_prime [Nontrivial R] + {p q : ℕ} [CharP R p] (hq : q.Prime) (hneq : p ≠ q) : (q : R) ≠ 0 := fun h ↦ by + rw [cast_eq_zero_iff R p q] at h + rcases hq.eq_one_or_self_of_dvd _ h with h | h + · subst h + exact false_of_nontrivial_of_char_one (R := R) + · exact hneq h + +lemma ringChar_of_prime_eq_zero [Nontrivial R] {p : ℕ} (hprime : Nat.Prime p) + (hp0 : (p : R) = 0) : ringChar R = p := + Or.resolve_left ((Nat.dvd_prime hprime).1 (ringChar.dvd hp0)) ringChar_ne_one + +lemma charP_iff_prime_eq_zero [Nontrivial R] {p : ℕ} (hp : p.Prime) : + CharP R p ↔ (p : R) = 0 := + ⟨fun _ => cast_eq_zero R p, + fun hp0 => (ringChar_of_prime_eq_zero hp hp0) ▸ inferInstance⟩ + +end NonAssocSemiring end CharP -end CommRing - - -namespace CharP section -variable (R) [NonAssocRing R] - -/-- The characteristic of a finite ring cannot be zero. -/ -theorem char_ne_zero_of_finite (p : ℕ) [CharP R p] [Finite R] : p ≠ 0 := by - rintro rfl - haveI : CharZero R := charP_to_charZero R - cases nonempty_fintype R - exact absurd Nat.cast_injective (not_injective_infinite_finite ((↑) : ℕ → R)) - -theorem ringChar_ne_zero_of_finite [Finite R] : ringChar R ≠ 0 := - char_ne_zero_of_finite R (ringChar R) +/-- We have `2 ≠ 0` in a nontrivial ring whose characteristic is not `2`. -/ +protected lemma Ring.two_ne_zero {R : Type*} [NonAssocSemiring R] [Nontrivial R] + (hR : ringChar R ≠ 2) : (2 : R) ≠ 0 := by + rw [Ne, (by norm_cast : (2 : R) = (2 : ℕ)), ringChar.spec, Nat.dvd_prime Nat.prime_two] + exact mt (or_iff_left hR).mp CharP.ringChar_ne_one + +-- We have `CharP.neg_one_ne_one`, which assumes `[Ring R] (p : ℕ) [CharP R p] [Fact (2 < p)]`. +-- This is a version using `ringChar` instead. +/-- Characteristic `≠ 2` and nontrivial implies that `-1 ≠ 1`. -/ +lemma Ring.neg_one_ne_one_of_char_ne_two {R : Type*} [NonAssocRing R] [Nontrivial R] + (hR : ringChar R ≠ 2) : (-1 : R) ≠ 1 := fun h => + Ring.two_ne_zero hR (one_add_one_eq_two (R := R) ▸ neg_eq_iff_add_eq_zero.mp h) + +/-- Characteristic `≠ 2` in a domain implies that `-a = a` iff `a = 0`. -/ +lemma Ring.eq_self_iff_eq_zero_of_char_ne_two {R : Type*} [NonAssocRing R] [Nontrivial R] + [NoZeroDivisors R] (hR : ringChar R ≠ 2) {a : R} : -a = a ↔ a = 0 := + ⟨fun h => + (mul_eq_zero.mp <| (two_mul a).trans <| neg_eq_iff_add_eq_zero.mp h).resolve_left + (Ring.two_ne_zero hR), + fun h => ((congr_arg (fun x => -x) h).trans neg_zero).trans h.symm⟩ end -section Ring - -variable (R) [Ring R] [NoZeroDivisors R] [Nontrivial R] [Finite R] - -theorem char_is_prime (p : ℕ) [CharP R p] : p.Prime := - Or.resolve_right (char_is_prime_or_zero R p) (char_ne_zero_of_finite R p) - -end Ring -end CharP - -/-! ### The Frobenius automorphism -/ - -section frobenius -section CommSemiring -variable [CommSemiring R] {S : Type*} [CommSemiring S] (f : R →* S) (g : R →+* S) (p m n : ℕ) - [ExpChar R p] [ExpChar S p] (x y : R) +section Prod +variable (S : Type*) [AddMonoidWithOne R] [AddMonoidWithOne S] (p q : ℕ) [CharP R p] -open ExpChar +/-- The characteristic of the product of rings is the least common multiple of the +characteristics of the two rings. -/ +instance Nat.lcm.charP [CharP S q] : CharP (R × S) (Nat.lcm p q) where + cast_eq_zero_iff' := by + simp [Prod.ext_iff, CharP.cast_eq_zero_iff R p, CharP.cast_eq_zero_iff S q, Nat.lcm_dvd_iff] -variable (R) in -/-- The frobenius map `x ↦ x ^ p`. -/ -def frobenius : R →+* R where - __ := powMonoidHom p - map_zero' := zero_pow (expChar_pos R p).ne' - map_add' _ _ := add_pow_expChar .. - -variable (R) in -/-- The iterated frobenius map `x ↦ x ^ p ^ n`. -/ -def iterateFrobenius : R →+* R where - __ := powMonoidHom (p ^ n) - map_zero' := zero_pow (expChar_pow_pos R p n).ne' - map_add' _ _ := add_pow_expChar_pow .. - - -lemma frobenius_def : frobenius R p x = x ^ p := rfl - -lemma iterateFrobenius_def : iterateFrobenius R p n x = x ^ p ^ n := rfl - -lemma iterate_frobenius : (frobenius R p)^[n] x = x ^ p ^ n := congr_fun (pow_iterate p n) x - -variable (R) - -lemma coe_iterateFrobenius : iterateFrobenius R p n = (frobenius R p)^[n] := - (pow_iterate p n).symm - -lemma iterateFrobenius_one_apply : iterateFrobenius R p 1 x = x ^ p := by - rw [iterateFrobenius_def, pow_one] - -@[simp] -lemma iterateFrobenius_one : iterateFrobenius R p 1 = frobenius R p := - RingHom.ext (iterateFrobenius_one_apply R p) +/-- The characteristic of the product of two rings of the same characteristic + is the same as the characteristic of the rings -/ +instance Prod.charP [CharP S p] : CharP (R × S) p := by + convert Nat.lcm.charP R S p p; simp -lemma iterateFrobenius_zero_apply : iterateFrobenius R p 0 x = x := by - rw [iterateFrobenius_def, pow_zero, pow_one] +instance Prod.charZero_of_left [CharZero R] : CharZero (R × S) where + cast_injective _ _ h := CharZero.cast_injective congr(Prod.fst $h) -@[simp] -lemma iterateFrobenius_zero : iterateFrobenius R p 0 = RingHom.id R := - RingHom.ext (iterateFrobenius_zero_apply R p) +instance Prod.charZero_of_right [CharZero S] : CharZero (R × S) where + cast_injective _ _ h := CharZero.cast_injective congr(Prod.snd $h) -lemma iterateFrobenius_add_apply : - iterateFrobenius R p (m + n) x = iterateFrobenius R p m (iterateFrobenius R p n x) := by - simp_rw [iterateFrobenius_def, add_comm m n, pow_add, pow_mul] +end Prod -lemma iterateFrobenius_add : - iterateFrobenius R p (m + n) = (iterateFrobenius R p m).comp (iterateFrobenius R p n) := - RingHom.ext (iterateFrobenius_add_apply R p m n) +instance ULift.charP [AddMonoidWithOne R] (p : ℕ) [CharP R p] : CharP (ULift R) p where + cast_eq_zero_iff' n := Iff.trans ULift.ext_iff <| CharP.cast_eq_zero_iff R p n -lemma iterateFrobenius_mul_apply : - iterateFrobenius R p (m * n) x = (iterateFrobenius R p m)^[n] x := by - simp_rw [coe_iterateFrobenius, Function.iterate_mul] +instance MulOpposite.charP [AddMonoidWithOne R] (p : ℕ) [CharP R p] : CharP Rᵐᵒᵖ p where + cast_eq_zero_iff' n := MulOpposite.unop_inj.symm.trans <| CharP.cast_eq_zero_iff R p n -lemma coe_iterateFrobenius_mul : iterateFrobenius R p (m * n) = (iterateFrobenius R p m)^[n] := - funext (iterateFrobenius_mul_apply R p m n) - -variable {R} - -lemma frobenius_mul : frobenius R p (x * y) = frobenius R p x * frobenius R p y := - map_mul (frobenius R p) x y - -lemma frobenius_one : frobenius R p 1 = 1 := one_pow _ - -lemma MonoidHom.map_frobenius : f (frobenius R p x) = frobenius S p (f x) := map_pow f x p -lemma RingHom.map_frobenius : g (frobenius R p x) = frobenius S p (g x) := map_pow g x p - -lemma MonoidHom.map_iterate_frobenius (n : ℕ) : - f ((frobenius R p)^[n] x) = (frobenius S p)^[n] (f x) := - Function.Semiconj.iterate_right (f.map_frobenius p) n x - -lemma RingHom.map_iterate_frobenius (n : ℕ) : - g ((frobenius R p)^[n] x) = (frobenius S p)^[n] (g x) := - g.toMonoidHom.map_iterate_frobenius p x n - -lemma MonoidHom.iterate_map_frobenius (f : R →* R) (p : ℕ) [ExpChar R p] (n : ℕ) : - f^[n] (frobenius R p x) = frobenius R p (f^[n] x) := - iterate_map_pow f _ _ _ - -lemma RingHom.iterate_map_frobenius (f : R →+* R) (p : ℕ) [ExpChar R p] (n : ℕ) : - f^[n] (frobenius R p x) = frobenius R p (f^[n] x) := iterate_map_pow f _ _ _ - -lemma list_sum_pow_char (l : List R) : l.sum ^ p = (l.map (· ^ p : R → R)).sum := - map_list_sum (frobenius R p) _ +section -lemma multiset_sum_pow_char (s : Multiset R) : s.sum ^ p = (s.map (· ^ p : R → R)).sum := - map_multiset_sum (frobenius R p) _ +/-- If two integers from `{0, 1, -1}` result in equal elements in a ring `R` +that is nontrivial and of characteristic not `2`, then they are equal. -/ +lemma Int.cast_injOn_of_ringChar_ne_two {R : Type*} [NonAssocRing R] [Nontrivial R] + (hR : ringChar R ≠ 2) : ({0, 1, -1} : Set ℤ).InjOn ((↑) : ℤ → R) := by + rintro _ (rfl | rfl | rfl) _ (rfl | rfl | rfl) h <;> + simp only + [cast_neg, cast_one, cast_zero, neg_eq_zero, one_ne_zero, zero_ne_one, zero_eq_neg] at h ⊢ + · exact ((Ring.neg_one_ne_one_of_char_ne_two hR).symm h).elim + · exact ((Ring.neg_one_ne_one_of_char_ne_two hR) h).elim -lemma sum_pow_char {ι : Type*} (s : Finset ι) (f : ι → R) : (∑ i ∈ s, f i) ^ p = ∑ i ∈ s, f i ^ p := - map_sum (frobenius R p) _ _ +end -variable (n : ℕ) +namespace CharZero -lemma list_sum_pow_char_pow (l : List R) : l.sum ^ p ^ n = (l.map (· ^ p ^ n : R → R)).sum := - map_list_sum (iterateFrobenius R p n) _ +lemma charZero_iff_forall_prime_ne_zero [NonAssocRing R] [NoZeroDivisors R] [Nontrivial R] : + CharZero R ↔ ∀ p : ℕ, p.Prime → (p : R) ≠ 0 := by + refine ⟨fun h p hp => by simp [hp.ne_zero], fun h => ?_⟩ + let p := ringChar R + cases CharP.char_is_prime_or_zero R p with + | inl hp => simpa using h p hp + | inr h => have : CharP R 0 := h ▸ inferInstance; exact CharP.charP_to_charZero R -lemma multiset_sum_pow_char_pow (s : Multiset R) : - s.sum ^ p ^ n = (s.map (· ^ p ^ n : R → R)).sum := map_multiset_sum (iterateFrobenius R p n) _ +end CharZero -lemma sum_pow_char_pow {ι : Type*} (s : Finset ι) (f : ι → R) : - (∑ i ∈ s, f i) ^ p ^ n = ∑ i ∈ s, f i ^ p ^ n := map_sum (iterateFrobenius R p n) _ _ +namespace Fin -end CommSemiring +instance charP (n : ℕ) [NeZero n] : CharP (Fin n) n where cast_eq_zero_iff' _ := natCast_eq_zero -section CommRing -variable [CommRing R] (p : ℕ) [ExpChar R p] (x y : R) +end Fin -lemma frobenius_neg : frobenius R p (-x) = -frobenius R p x := map_neg .. +section AddMonoidWithOne +variable [AddMonoidWithOne R] -lemma frobenius_sub : frobenius R p (x - y) = frobenius R p x - frobenius R p y := map_sub .. +instance (S : Type*) [Semiring S] (p) [ExpChar R p] [ExpChar S p] : ExpChar (R × S) p := by + obtain hp | ⟨hp⟩ := ‹ExpChar R p› + · have := Prod.charZero_of_left R S; exact .zero + obtain _ | _ := ‹ExpChar S p› + · exact (Nat.not_prime_one hp).elim + · have := Prod.charP R S p; exact .prime hp -end CommRing -end frobenius +end AddMonoidWithOne diff --git a/Mathlib/Algebra/CharP/CharAndCard.lean b/Mathlib/Algebra/CharP/CharAndCard.lean index 75671724d2c94..f78ed1df476fc 100644 --- a/Mathlib/Algebra/CharP/CharAndCard.lean +++ b/Mathlib/Algebra/CharP/CharAndCard.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Michael Stoll -/ import Mathlib.Algebra.CharP.Basic +import Mathlib.Algebra.CharP.Lemmas import Mathlib.GroupTheory.Perm.Cycle.Type import Mathlib.RingTheory.Coprime.Lemmas diff --git a/Mathlib/Algebra/CharP/Defs.lean b/Mathlib/Algebra/CharP/Defs.lean index 9509007537ca2..a80dd345ea9ac 100644 --- a/Mathlib/Algebra/CharP/Defs.lean +++ b/Mathlib/Algebra/CharP/Defs.lean @@ -3,19 +3,21 @@ Copyright (c) 2018 Kenny Lau. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau, Joey van Langen, Casper Putz -/ -import Mathlib.Algebra.Field.Basic -import Mathlib.Algebra.Group.Fin.Basic -import Mathlib.Algebra.Group.ULift import Mathlib.Data.Int.ModEq -import Mathlib.Data.Nat.Cast.Prod +import Mathlib.Data.Nat.Cast.Defs import Mathlib.Data.Nat.Find import Mathlib.Data.Nat.Prime.Defs -import Mathlib.Data.ULift import Mathlib.Tactic.NormNum.Basic -import Mathlib.Order.Interval.Set.Basic /-! # Characteristic of semirings + +## Main definitions + * `CharP R p` expresses that the ring (additive monoid with one) `R` has characteristic `p` + * `ringChar`: the characteristic of a ring + * `ExpChar R p` expresses that the ring (additive monoid with one) `R` has + exponential characteristic `p` (which is `1` if `R` has characteristic 0, and `p` if it has + prime characteristic `p`) -/ assert_not_exists Finset @@ -87,21 +89,11 @@ lemma natCast_eq_natCast : (a : R) = b ↔ a ≡ b [MOD p] := by ← add_right_cancel_iff (G := R) (a := a) (b := b - a), zero_add, ← Nat.cast_add, Nat.sub_add_cancel hle, eq_comm] -lemma natCast_injOn_Iio : (Set.Iio p).InjOn ((↑) : ℕ → R) := - fun _a ha _b hb hab ↦ ((natCast_eq_natCast _ _).1 hab).eq_of_lt_of_lt ha hb - end AddMonoidWithOne section AddGroupWithOne variable [AddGroupWithOne R] (p : ℕ) [CharP R p] {a b : ℤ} -lemma intCast_injOn_Ico [IsRightCancelAdd R] : InjOn (Int.cast : ℤ → R) (Ico 0 p) := by - rintro a ⟨ha₀, ha⟩ b ⟨hb₀, hb⟩ hab - lift a to ℕ using ha₀ - lift b to ℕ using hb₀ - norm_cast at * - exact natCast_injOn_Iio _ _ ha hb hab - lemma intCast_eq_zero_iff (a : ℤ) : (a : R) = 0 ↔ (p : ℤ) ∣ a := by rcases lt_trichotomy a 0 with (h | rfl | h) · rw [← neg_eq_zero, ← Int.cast_neg, ← Int.dvd_neg] @@ -212,10 +204,6 @@ lemma CharP.neg_one_ne_one [Ring R] (p : ℕ) [CharP R p] [Fact (2 < p)] : (-1 : rw [fact_iff] at * omega -lemma RingHom.charP_iff_charP {K L : Type*} [DivisionRing K] [Semiring L] [Nontrivial L] - (f : K →+* L) (p : ℕ) : CharP K p ↔ CharP L p := by - simp only [charP_iff, ← f.injective.eq_iff, map_natCast f, map_zero f] - namespace CharP section @@ -311,17 +299,6 @@ lemma false_of_nontrivial_of_char_one [Nontrivial R] [CharP R 1] : False := by have : Subsingleton R := CharOne.subsingleton exact false_of_nontrivial_of_subsingleton R -variable (R) in -/-- If a ring `R` is of characteristic `p`, then for any prime number `q` different from `p`, -it is not zero in `R`. -/ -lemma cast_ne_zero_of_ne_of_prime [Nontrivial R] - {p q : ℕ} [CharP R p] (hq : q.Prime) (hneq : p ≠ q) : (q : R) ≠ 0 := fun h ↦ by - rw [cast_eq_zero_iff R p q] at h - rcases hq.eq_one_or_self_of_dvd _ h with h | h - · subst h - exact false_of_nontrivial_of_char_one (R := R) - · exact hneq h - lemma ringChar_ne_one [Nontrivial R] : ringChar R ≠ 1 := by intro h apply zero_ne_one' R @@ -332,85 +309,9 @@ lemma nontrivial_of_char_ne_one {v : ℕ} (hv : v ≠ 1) [hr : CharP R v] : Nont ⟨⟨(1 : ℕ), 0, fun h => hv <| by rwa [CharP.cast_eq_zero_iff _ v, Nat.dvd_one] at h⟩⟩ -lemma ringChar_of_prime_eq_zero [Nontrivial R] {p : ℕ} (hprime : Nat.Prime p) - (hp0 : (p : R) = 0) : ringChar R = p := - Or.resolve_left ((Nat.dvd_prime hprime).1 (ringChar.dvd hp0)) ringChar_ne_one - -lemma charP_iff_prime_eq_zero [Nontrivial R] {p : ℕ} (hp : p.Prime) : - CharP R p ↔ (p : R) = 0 := - ⟨fun _ => cast_eq_zero R p, - fun hp0 => (ringChar_of_prime_eq_zero hp hp0) ▸ inferInstance⟩ - end NonAssocSemiring end CharP -section - -/-- We have `2 ≠ 0` in a nontrivial ring whose characteristic is not `2`. -/ -protected lemma Ring.two_ne_zero {R : Type*} [NonAssocSemiring R] [Nontrivial R] - (hR : ringChar R ≠ 2) : (2 : R) ≠ 0 := by - rw [Ne, (by norm_cast : (2 : R) = (2 : ℕ)), ringChar.spec, Nat.dvd_prime Nat.prime_two] - exact mt (or_iff_left hR).mp CharP.ringChar_ne_one - --- We have `CharP.neg_one_ne_one`, which assumes `[Ring R] (p : ℕ) [CharP R p] [Fact (2 < p)]`. --- This is a version using `ringChar` instead. -/-- Characteristic `≠ 2` and nontrivial implies that `-1 ≠ 1`. -/ -lemma Ring.neg_one_ne_one_of_char_ne_two {R : Type*} [NonAssocRing R] [Nontrivial R] - (hR : ringChar R ≠ 2) : (-1 : R) ≠ 1 := fun h => - Ring.two_ne_zero hR (one_add_one_eq_two (R := R) ▸ neg_eq_iff_add_eq_zero.mp h) - -/-- Characteristic `≠ 2` in a domain implies that `-a = a` iff `a = 0`. -/ -lemma Ring.eq_self_iff_eq_zero_of_char_ne_two {R : Type*} [NonAssocRing R] [Nontrivial R] - [NoZeroDivisors R] (hR : ringChar R ≠ 2) {a : R} : -a = a ↔ a = 0 := - ⟨fun h => - (mul_eq_zero.mp <| (two_mul a).trans <| neg_eq_iff_add_eq_zero.mp h).resolve_left - (Ring.two_ne_zero hR), - fun h => ((congr_arg (fun x => -x) h).trans neg_zero).trans h.symm⟩ - -end - -section Prod -variable (S : Type*) [AddMonoidWithOne R] [AddMonoidWithOne S] (p q : ℕ) [CharP R p] - -/-- The characteristic of the product of rings is the least common multiple of the -characteristics of the two rings. -/ -instance Nat.lcm.charP [CharP S q] : CharP (R × S) (Nat.lcm p q) where - cast_eq_zero_iff' := by - simp [Prod.ext_iff, CharP.cast_eq_zero_iff R p, CharP.cast_eq_zero_iff S q, Nat.lcm_dvd_iff] - -/-- The characteristic of the product of two rings of the same characteristic - is the same as the characteristic of the rings -/ -instance Prod.charP [CharP S p] : CharP (R × S) p := by - convert Nat.lcm.charP R S p p; simp - -instance Prod.charZero_of_left [CharZero R] : CharZero (R × S) where - cast_injective _ _ h := CharZero.cast_injective congr(Prod.fst $h) - -instance Prod.charZero_of_right [CharZero S] : CharZero (R × S) where - cast_injective _ _ h := CharZero.cast_injective congr(Prod.snd $h) - -end Prod - -instance ULift.charP [AddMonoidWithOne R] (p : ℕ) [CharP R p] : CharP (ULift R) p where - cast_eq_zero_iff' n := Iff.trans ULift.ext_iff <| CharP.cast_eq_zero_iff R p n - -instance MulOpposite.charP [AddMonoidWithOne R] (p : ℕ) [CharP R p] : CharP Rᵐᵒᵖ p where - cast_eq_zero_iff' n := MulOpposite.unop_inj.symm.trans <| CharP.cast_eq_zero_iff R p n - -section - -/-- If two integers from `{0, 1, -1}` result in equal elements in a ring `R` -that is nontrivial and of characteristic not `2`, then they are equal. -/ -lemma Int.cast_injOn_of_ringChar_ne_two {R : Type*} [NonAssocRing R] [Nontrivial R] - (hR : ringChar R ≠ 2) : ({0, 1, -1} : Set ℤ).InjOn ((↑) : ℤ → R) := by - rintro _ (rfl | rfl | rfl) _ (rfl | rfl | rfl) h <;> - simp only - [cast_neg, cast_one, cast_zero, neg_eq_zero, one_ne_zero, zero_ne_one, zero_eq_neg] at h ⊢ - · exact ((Ring.neg_one_ne_one_of_char_ne_two hR).symm h).elim - · exact ((Ring.neg_one_ne_one_of_char_ne_two hR) h).elim - -end - namespace NeZero variable [AddMonoidWithOne R] {r : R} {n p : ℕ} @@ -423,24 +324,6 @@ lemma not_char_dvd (p : ℕ) [CharP R p] (k : ℕ) [h : NeZero (k : R)] : ¬p end NeZero -namespace CharZero - -lemma charZero_iff_forall_prime_ne_zero [NonAssocRing R] [NoZeroDivisors R] [Nontrivial R] : - CharZero R ↔ ∀ p : ℕ, p.Prime → (p : R) ≠ 0 := by - refine ⟨fun h p hp => by simp [hp.ne_zero], fun h => ?_⟩ - let p := ringChar R - cases CharP.char_is_prime_or_zero R p with - | inl hp => simpa using h p hp - | inr h => have : CharP R 0 := h ▸ inferInstance; exact CharP.charP_to_charZero R - -end CharZero - -namespace Fin - -instance charP (n : ℕ) [NeZero n] : CharP (Fin n) n where cast_eq_zero_iff' _ := natCast_eq_zero - -end Fin - /-! ### Exponential characteristic @@ -469,13 +352,6 @@ lemma expChar_ne_zero (p : ℕ) [hR : ExpChar R p] : p ≠ 0 := by · exact one_ne_zero · exact ‹p.Prime›.ne_zero -instance (S : Type*) [Semiring S] (p) [ExpChar R p] [ExpChar S p] : ExpChar (R × S) p := by - obtain hp | ⟨hp⟩ := ‹ExpChar R p› - · have := Prod.charZero_of_left R S; exact .zero - obtain _ | _ := ‹ExpChar S p› - · exact (Nat.not_prime_one hp).elim - · have := Prod.charP R S p; exact .prime hp - variable {R} in /-- The exponential characteristic is unique. -/ lemma ExpChar.eq {p q : ℕ} (hp : ExpChar R p) (hq : ExpChar R q) : p = q := by diff --git a/Mathlib/Algebra/CharP/ExpChar.lean b/Mathlib/Algebra/CharP/ExpChar.lean index ebc8479fc69c9..9053124e1a464 100644 --- a/Mathlib/Algebra/CharP/ExpChar.lean +++ b/Mathlib/Algebra/CharP/ExpChar.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jakob Scholbach -/ import Mathlib.Algebra.Algebra.Defs -import Mathlib.Algebra.CharP.Basic +import Mathlib.Algebra.CharP.Lemmas /-! # Exponential characteristic diff --git a/Mathlib/Algebra/CharP/Lemmas.lean b/Mathlib/Algebra/CharP/Lemmas.lean new file mode 100644 index 0000000000000..271aa60a6d144 --- /dev/null +++ b/Mathlib/Algebra/CharP/Lemmas.lean @@ -0,0 +1,391 @@ +/- +Copyright (c) 2018 Kenny Lau. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kenny Lau, Joey van Langen, Casper Putz +-/ +import Mathlib.Algebra.CharP.Defs +import Mathlib.Algebra.GroupPower.IterateHom +import Mathlib.Data.Nat.Multiplicity +import Mathlib.Data.Nat.Choose.Sum + +/-! +# Characteristic of semirings +-/ + +assert_not_exists Algebra +assert_not_exists LinearMap +assert_not_exists orderOf + +open Finset + +variable {R S : Type*} + +namespace Commute + +variable [Semiring R] {p : ℕ} (hp : p.Prime) {x y : R} +include hp + +protected theorem add_pow_prime_pow_eq (h : Commute x y) (n : ℕ) : + (x + y) ^ p ^ n = + x ^ p ^ n + y ^ p ^ n + + p * ∑ k ∈ Ioo 0 (p ^ n), x ^ k * y ^ (p ^ n - k) * ↑((p ^ n).choose k / p) := by + trans x ^ p ^ n + y ^ p ^ n + ∑ k ∈ Ioo 0 (p ^ n), x ^ k * y ^ (p ^ n - k) * (p ^ n).choose k + · simp_rw [h.add_pow, ← Nat.Ico_zero_eq_range, Nat.Ico_succ_right, Icc_eq_cons_Ico (zero_le _), + Finset.sum_cons, Ico_eq_cons_Ioo (pow_pos hp.pos _), Finset.sum_cons, tsub_self, tsub_zero, + pow_zero, Nat.choose_zero_right, Nat.choose_self, Nat.cast_one, mul_one, one_mul, ← add_assoc] + · congr 1 + simp_rw [Finset.mul_sum, Nat.cast_comm, mul_assoc _ _ (p : R), ← Nat.cast_mul] + refine Finset.sum_congr rfl fun i hi => ?_ + rw [mem_Ioo] at hi + rw [Nat.div_mul_cancel (hp.dvd_choose_pow hi.1.ne' hi.2.ne)] + +protected theorem add_pow_prime_eq (h : Commute x y) : + (x + y) ^ p = + x ^ p + y ^ p + p * ∑ k ∈ Finset.Ioo 0 p, x ^ k * y ^ (p - k) * ↑(p.choose k / p) := by + simpa using h.add_pow_prime_pow_eq hp 1 + +protected theorem exists_add_pow_prime_pow_eq (h : Commute x y) (n : ℕ) : + ∃ r, (x + y) ^ p ^ n = x ^ p ^ n + y ^ p ^ n + p * r := + ⟨_, h.add_pow_prime_pow_eq hp n⟩ + +protected theorem exists_add_pow_prime_eq (h : Commute x y) : + ∃ r, (x + y) ^ p = x ^ p + y ^ p + p * r := + ⟨_, h.add_pow_prime_eq hp⟩ + +end Commute + +section CommSemiring + +variable [CommSemiring R] {p : ℕ} (hp : p.Prime) (x y : R) (n : ℕ) +include hp + +theorem add_pow_prime_pow_eq : + (x + y) ^ p ^ n = + x ^ p ^ n + y ^ p ^ n + + p * ∑ k ∈ Finset.Ioo 0 (p ^ n), x ^ k * y ^ (p ^ n - k) * ↑((p ^ n).choose k / p) := + (Commute.all x y).add_pow_prime_pow_eq hp n + +theorem add_pow_prime_eq : + (x + y) ^ p = + x ^ p + y ^ p + p * ∑ k ∈ Finset.Ioo 0 p, x ^ k * y ^ (p - k) * ↑(p.choose k / p) := + (Commute.all x y).add_pow_prime_eq hp + +theorem exists_add_pow_prime_pow_eq : + ∃ r, (x + y) ^ p ^ n = x ^ p ^ n + y ^ p ^ n + p * r := + (Commute.all x y).exists_add_pow_prime_pow_eq hp n + +theorem exists_add_pow_prime_eq : + ∃ r, (x + y) ^ p = x ^ p + y ^ p + p * r := + (Commute.all x y).exists_add_pow_prime_eq hp + +end CommSemiring + +section Semiring +variable [Semiring R] {x y : R} (p n : ℕ) + +section ExpChar +variable [hR : ExpChar R p] + +lemma add_pow_expChar_of_commute (h : Commute x y) : (x + y) ^ p = x ^ p + y ^ p := by + obtain _ | hprime := hR + · simp only [pow_one] + · let ⟨r, hr⟩ := h.exists_add_pow_prime_eq hprime + simp [hr] + +lemma add_pow_expChar_pow_of_commute (h : Commute x y) : + (x + y) ^ p ^ n = x ^ p ^ n + y ^ p ^ n := by + obtain _ | hprime := hR + · simp only [one_pow, pow_one] + · let ⟨r, hr⟩ := h.exists_add_pow_prime_pow_eq hprime n + simp [hr] + +lemma add_pow_eq_mul_pow_add_pow_div_expChar_of_commute (h : Commute x y) : + (x + y) ^ n = (x + y) ^ (n % p) * (x ^ p + y ^ p) ^ (n / p) := by + rw [← add_pow_expChar_of_commute _ h, ← pow_mul, ← pow_add, Nat.mod_add_div] + +end ExpChar + +section CharP +variable [hp : Fact p.Prime] [CharP R p] + +lemma add_pow_char_of_commute (h : Commute x y) : (x + y) ^ p = x ^ p + y ^ p := + add_pow_expChar_of_commute _ h + +lemma add_pow_char_pow_of_commute (h : Commute x y) : (x + y) ^ p ^ n = x ^ p ^ n + y ^ p ^ n := + add_pow_expChar_pow_of_commute _ _ h + +lemma add_pow_eq_mul_pow_add_pow_div_char_of_commute (h : Commute x y) : + (x + y) ^ n = (x + y) ^ (n % p) * (x ^ p + y ^ p) ^ (n / p) := + add_pow_eq_mul_pow_add_pow_div_expChar_of_commute _ _ h + +end CharP +end Semiring + +section CommSemiring +variable [CommSemiring R] (x y : R) (p n : ℕ) + +section ExpChar +variable [hR : ExpChar R p] + +lemma add_pow_expChar : (x + y) ^ p = x ^ p + y ^ p := add_pow_expChar_of_commute _ <| .all .. + +lemma add_pow_expChar_pow : (x + y) ^ p ^ n = x ^ p ^ n + y ^ p ^ n := + add_pow_expChar_pow_of_commute _ _ <| .all .. + +lemma add_pow_eq_mul_pow_add_pow_div_expChar : + (x + y) ^ n = (x + y) ^ (n % p) * (x ^ p + y ^ p) ^ (n / p) := + add_pow_eq_mul_pow_add_pow_div_expChar_of_commute _ _ <| .all .. + +end ExpChar + +section CharP +variable [hp : Fact p.Prime] [CharP R p] + +lemma add_pow_char : (x + y) ^ p = x ^ p + y ^ p := add_pow_expChar .. +lemma add_pow_char_pow : (x + y) ^ p ^ n = x ^ p ^ n + y ^ p ^ n := add_pow_expChar_pow .. + +lemma add_pow_eq_mul_pow_add_pow_div_char : + (x + y) ^ n = (x + y) ^ (n % p) * (x ^ p + y ^ p) ^ (n / p) := + add_pow_eq_mul_pow_add_pow_div_expChar .. + +@[deprecated (since := "2024-10-21")] +alias add_pow_eq_add_pow_mod_mul_pow_add_pow_div := add_pow_eq_mul_pow_add_pow_div_char + +end CharP +end CommSemiring + +section Ring +variable [Ring R] {x y : R} (p n : ℕ) + +section ExpChar +variable [hR : ExpChar R p] +include hR + +lemma sub_pow_expChar_of_commute (h : Commute x y) : (x - y) ^ p = x ^ p - y ^ p := by + simp [eq_sub_iff_add_eq, ← add_pow_expChar_of_commute _ (h.sub_left rfl)] + +lemma sub_pow_expChar_pow_of_commute (h : Commute x y) : + (x - y) ^ p ^ n = x ^ p ^ n - y ^ p ^ n := by + simp [eq_sub_iff_add_eq, ← add_pow_expChar_pow_of_commute _ _ (h.sub_left rfl)] + +lemma sub_pow_eq_mul_pow_sub_pow_div_expChar_of_commute (h : Commute x y) : + (x - y) ^ n = (x - y) ^ (n % p) * (x ^ p - y ^ p) ^ (n / p) := by + rw [← sub_pow_expChar_of_commute _ h, ← pow_mul, ← pow_add, Nat.mod_add_div] + +variable (R) + +lemma neg_one_pow_expChar : (-1 : R) ^ p = -1 := by + rw [eq_neg_iff_add_eq_zero] + nth_rw 2 [← one_pow p] + rw [← add_pow_expChar_of_commute _ (Commute.one_right _), neg_add_cancel, + zero_pow (expChar_ne_zero R p)] + +lemma neg_one_pow_expChar_pow : (-1 : R) ^ p ^ n = -1 := by + rw [eq_neg_iff_add_eq_zero] + nth_rw 2 [← one_pow (p ^ n)] + rw [← add_pow_expChar_pow_of_commute _ _ (Commute.one_right _), neg_add_cancel, + zero_pow (pow_ne_zero _ <| expChar_ne_zero R p)] + +end ExpChar + +section CharP +variable [hp : Fact p.Prime] [CharP R p] + +lemma sub_pow_char_of_commute (h : Commute x y) : (x - y) ^ p = x ^ p - y ^ p := + sub_pow_expChar_of_commute _ h + +lemma sub_pow_char_pow_of_commute (h : Commute x y) : (x - y) ^ p ^ n = x ^ p ^ n - y ^ p ^ n := + sub_pow_expChar_pow_of_commute _ _ h + +variable (R) + +lemma neg_one_pow_char : (-1 : R) ^ p = -1 := neg_one_pow_expChar .. + +lemma neg_one_pow_char_pow : (-1 : R) ^ p ^ n = -1 := neg_one_pow_expChar_pow .. + +lemma sub_pow_eq_mul_pow_sub_pow_div_char_of_commute (h : Commute x y) : + (x - y) ^ n = (x - y) ^ (n % p) * (x ^ p - y ^ p) ^ (n / p) := + sub_pow_eq_mul_pow_sub_pow_div_expChar_of_commute _ _ h + +end CharP +end Ring + +section CommRing +variable [CommRing R] (x y : R) (n : ℕ) {p : ℕ} + +section ExpChar +variable [hR : ExpChar R p] + +lemma sub_pow_expChar : (x - y) ^ p = x ^ p - y ^ p := sub_pow_expChar_of_commute _ <| .all .. + +lemma sub_pow_expChar_pow : (x - y) ^ p ^ n = x ^ p ^ n - y ^ p ^ n := + sub_pow_expChar_pow_of_commute _ _ <| .all .. + +lemma sub_pow_eq_mul_pow_sub_pow_div_expChar : + (x - y) ^ n = (x - y) ^ (n % p) * (x ^ p - y ^ p) ^ (n / p) := + sub_pow_eq_mul_pow_sub_pow_div_expChar_of_commute _ _ <| .all .. + +end ExpChar + +section CharP +variable [hp : Fact p.Prime] [CharP R p] + +lemma sub_pow_char : (x - y) ^ p = x ^ p - y ^ p := sub_pow_expChar .. +lemma sub_pow_char_pow : (x - y) ^ p ^ n = x ^ p ^ n - y ^ p ^ n := sub_pow_expChar_pow .. + +lemma sub_pow_eq_mul_pow_sub_pow_div_char : + (x - y) ^ n = (x - y) ^ (n % p) * (x ^ p - y ^ p) ^ (n / p) := + sub_pow_eq_mul_pow_sub_pow_div_expChar .. + +end CharP +end CommRing + + +namespace CharP + +section + +variable (R) [NonAssocRing R] + +/-- The characteristic of a finite ring cannot be zero. -/ +theorem char_ne_zero_of_finite (p : ℕ) [CharP R p] [Finite R] : p ≠ 0 := by + rintro rfl + haveI : CharZero R := charP_to_charZero R + cases nonempty_fintype R + exact absurd Nat.cast_injective (not_injective_infinite_finite ((↑) : ℕ → R)) + +theorem ringChar_ne_zero_of_finite [Finite R] : ringChar R ≠ 0 := + char_ne_zero_of_finite R (ringChar R) + +end + +section Ring + +variable (R) [Ring R] [NoZeroDivisors R] [Nontrivial R] [Finite R] + +theorem char_is_prime (p : ℕ) [CharP R p] : p.Prime := + Or.resolve_right (char_is_prime_or_zero R p) (char_ne_zero_of_finite R p) + +end Ring +end CharP + +/-! ### The Frobenius automorphism -/ + +section frobenius +section CommSemiring +variable [CommSemiring R] {S : Type*} [CommSemiring S] (f : R →* S) (g : R →+* S) (p m n : ℕ) + [ExpChar R p] [ExpChar S p] (x y : R) + +open ExpChar + +variable (R) in +/-- The frobenius map `x ↦ x ^ p`. -/ +def frobenius : R →+* R where + __ := powMonoidHom p + map_zero' := zero_pow (expChar_pos R p).ne' + map_add' _ _ := add_pow_expChar .. + +variable (R) in +/-- The iterated frobenius map `x ↦ x ^ p ^ n`. -/ +def iterateFrobenius : R →+* R where + __ := powMonoidHom (p ^ n) + map_zero' := zero_pow (expChar_pow_pos R p n).ne' + map_add' _ _ := add_pow_expChar_pow .. + + +lemma frobenius_def : frobenius R p x = x ^ p := rfl + +lemma iterateFrobenius_def : iterateFrobenius R p n x = x ^ p ^ n := rfl + +lemma iterate_frobenius : (frobenius R p)^[n] x = x ^ p ^ n := congr_fun (pow_iterate p n) x + +variable (R) + +lemma coe_iterateFrobenius : iterateFrobenius R p n = (frobenius R p)^[n] := + (pow_iterate p n).symm + +lemma iterateFrobenius_one_apply : iterateFrobenius R p 1 x = x ^ p := by + rw [iterateFrobenius_def, pow_one] + +@[simp] +lemma iterateFrobenius_one : iterateFrobenius R p 1 = frobenius R p := + RingHom.ext (iterateFrobenius_one_apply R p) + +lemma iterateFrobenius_zero_apply : iterateFrobenius R p 0 x = x := by + rw [iterateFrobenius_def, pow_zero, pow_one] + +@[simp] +lemma iterateFrobenius_zero : iterateFrobenius R p 0 = RingHom.id R := + RingHom.ext (iterateFrobenius_zero_apply R p) + +lemma iterateFrobenius_add_apply : + iterateFrobenius R p (m + n) x = iterateFrobenius R p m (iterateFrobenius R p n x) := by + simp_rw [iterateFrobenius_def, add_comm m n, pow_add, pow_mul] + +lemma iterateFrobenius_add : + iterateFrobenius R p (m + n) = (iterateFrobenius R p m).comp (iterateFrobenius R p n) := + RingHom.ext (iterateFrobenius_add_apply R p m n) + +lemma iterateFrobenius_mul_apply : + iterateFrobenius R p (m * n) x = (iterateFrobenius R p m)^[n] x := by + simp_rw [coe_iterateFrobenius, Function.iterate_mul] + +lemma coe_iterateFrobenius_mul : iterateFrobenius R p (m * n) = (iterateFrobenius R p m)^[n] := + funext (iterateFrobenius_mul_apply R p m n) + +variable {R} + +lemma frobenius_mul : frobenius R p (x * y) = frobenius R p x * frobenius R p y := + map_mul (frobenius R p) x y + +lemma frobenius_one : frobenius R p 1 = 1 := one_pow _ + +lemma MonoidHom.map_frobenius : f (frobenius R p x) = frobenius S p (f x) := map_pow f x p +lemma RingHom.map_frobenius : g (frobenius R p x) = frobenius S p (g x) := map_pow g x p + +lemma MonoidHom.map_iterate_frobenius (n : ℕ) : + f ((frobenius R p)^[n] x) = (frobenius S p)^[n] (f x) := + Function.Semiconj.iterate_right (f.map_frobenius p) n x + +lemma RingHom.map_iterate_frobenius (n : ℕ) : + g ((frobenius R p)^[n] x) = (frobenius S p)^[n] (g x) := + g.toMonoidHom.map_iterate_frobenius p x n + +lemma MonoidHom.iterate_map_frobenius (f : R →* R) (p : ℕ) [ExpChar R p] (n : ℕ) : + f^[n] (frobenius R p x) = frobenius R p (f^[n] x) := + iterate_map_pow f _ _ _ + +lemma RingHom.iterate_map_frobenius (f : R →+* R) (p : ℕ) [ExpChar R p] (n : ℕ) : + f^[n] (frobenius R p x) = frobenius R p (f^[n] x) := iterate_map_pow f _ _ _ + +lemma list_sum_pow_char (l : List R) : l.sum ^ p = (l.map (· ^ p : R → R)).sum := + map_list_sum (frobenius R p) _ + +lemma multiset_sum_pow_char (s : Multiset R) : s.sum ^ p = (s.map (· ^ p : R → R)).sum := + map_multiset_sum (frobenius R p) _ + +lemma sum_pow_char {ι : Type*} (s : Finset ι) (f : ι → R) : (∑ i ∈ s, f i) ^ p = ∑ i ∈ s, f i ^ p := + map_sum (frobenius R p) _ _ + +variable (n : ℕ) + +lemma list_sum_pow_char_pow (l : List R) : l.sum ^ p ^ n = (l.map (· ^ p ^ n : R → R)).sum := + map_list_sum (iterateFrobenius R p n) _ + +lemma multiset_sum_pow_char_pow (s : Multiset R) : + s.sum ^ p ^ n = (s.map (· ^ p ^ n : R → R)).sum := map_multiset_sum (iterateFrobenius R p n) _ + +lemma sum_pow_char_pow {ι : Type*} (s : Finset ι) (f : ι → R) : + (∑ i ∈ s, f i) ^ p ^ n = ∑ i ∈ s, f i ^ p ^ n := map_sum (iterateFrobenius R p n) _ _ + +end CommSemiring + +section CommRing +variable [CommRing R] (p : ℕ) [ExpChar R p] (x y : R) + +lemma frobenius_neg : frobenius R p (-x) = -frobenius R p x := map_neg .. + +lemma frobenius_sub : frobenius R p (x - y) = frobenius R p x - frobenius R p y := map_sub .. + +end CommRing +end frobenius diff --git a/Mathlib/Algebra/CharP/Reduced.lean b/Mathlib/Algebra/CharP/Reduced.lean index 7328d585042a2..deea5f7b1eb53 100644 --- a/Mathlib/Algebra/CharP/Reduced.lean +++ b/Mathlib/Algebra/CharP/Reduced.lean @@ -3,7 +3,7 @@ Copyright (c) 2018 Kenny Lau. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau, Joey van Langen, Casper Putz -/ -import Mathlib.Algebra.CharP.Basic +import Mathlib.Algebra.CharP.Lemmas import Mathlib.RingTheory.Nilpotent.Defs /-! diff --git a/Mathlib/Algebra/CharP/Two.lean b/Mathlib/Algebra/CharP/Two.lean index edd4d620d33e7..508390fa0fa72 100644 --- a/Mathlib/Algebra/CharP/Two.lean +++ b/Mathlib/Algebra/CharP/Two.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Eric Wieser. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Eric Wieser -/ -import Mathlib.Algebra.CharP.Basic +import Mathlib.Algebra.CharP.Lemmas import Mathlib.GroupTheory.OrderOfElement /-! diff --git a/Mathlib/Algebra/Polynomial/Expand.lean b/Mathlib/Algebra/Polynomial/Expand.lean index 25e32f8c28e20..4b443cdf08aa9 100644 --- a/Mathlib/Algebra/Polynomial/Expand.lean +++ b/Mathlib/Algebra/Polynomial/Expand.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Kenny Lau. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau -/ -import Mathlib.Algebra.CharP.Basic +import Mathlib.Algebra.CharP.Lemmas import Mathlib.RingTheory.Polynomial.Basic /-! diff --git a/Mathlib/Combinatorics/Additive/FreimanHom.lean b/Mathlib/Combinatorics/Additive/FreimanHom.lean index dab82259694cf..70de5725f129a 100644 --- a/Mathlib/Combinatorics/Additive/FreimanHom.lean +++ b/Mathlib/Combinatorics/Additive/FreimanHom.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies, Bhavik Mehta -/ import Mathlib.Algebra.BigOperators.Ring -import Mathlib.Algebra.CharP.Defs +import Mathlib.Algebra.CharP.Basic import Mathlib.Algebra.Group.Pointwise.Set.Basic import Mathlib.Algebra.Group.Submonoid.Operations import Mathlib.Algebra.Order.BigOperators.Group.Multiset diff --git a/Mathlib/Data/Nat/Choose/Lucas.lean b/Mathlib/Data/Nat/Choose/Lucas.lean index 0ada97a2cd2dd..e3a167ba89664 100644 --- a/Mathlib/Data/Nat/Choose/Lucas.lean +++ b/Mathlib/Data/Nat/Choose/Lucas.lean @@ -3,7 +3,7 @@ Copyright (c) 2023 Gareth Ma. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Gareth Ma -/ -import Mathlib.Algebra.CharP.Basic +import Mathlib.Algebra.CharP.Lemmas import Mathlib.Data.ZMod.Basic import Mathlib.RingTheory.Polynomial.Basic diff --git a/Mathlib/Data/ZMod/Basic.lean b/Mathlib/Data/ZMod/Basic.lean index 4c03f03635c6f..0200eb971668e 100644 --- a/Mathlib/Data/ZMod/Basic.lean +++ b/Mathlib/Data/ZMod/Basic.lean @@ -3,6 +3,7 @@ Copyright (c) 2018 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes -/ +import Mathlib.Algebra.CharP.Basic import Mathlib.Algebra.Module.End import Mathlib.Algebra.Ring.Prod import Mathlib.GroupTheory.GroupAction.SubMulAction @@ -11,7 +12,6 @@ import Mathlib.Tactic.FinCases import Mathlib.Tactic.Linarith import Mathlib.Data.Fintype.Units - /-! # Integers mod `n` diff --git a/Mathlib/Data/ZMod/Defs.lean b/Mathlib/Data/ZMod/Defs.lean index 40ebde64104ee..97124f237c44e 100644 --- a/Mathlib/Data/ZMod/Defs.lean +++ b/Mathlib/Data/ZMod/Defs.lean @@ -16,14 +16,14 @@ This file provides the basic details of `ZMod n`, including its commutative ring ## Implementation details -This used to be inlined into `Data.ZMod.Basic`. This file imports `CharP.Basic`, which is an +This used to be inlined into `Data.ZMod.Basic`. This file imports `CharP.Lemmas`, which is an issue; all `CharP` instances create an `Algebra (ZMod p) R` instance; however, this instance may not be definitionally equal to other `Algebra` instances (for example, `GaloisField` also has an `Algebra` instance as it is defined as a `SplittingField`). The way to fix this is to use the forgetful inheritance pattern, and make `CharP` carry the data of what the `smul` should be (so for example, the `smul` on the `GaloisField` `CharP` instance should be equal to the `smul` from its `SplittingField` structure); there is only one possible `ZMod p` algebra for any `p`, so this -is not an issue mathematically. For this to be possible, however, we need `CharP.Basic` to be +is not an issue mathematically. For this to be possible, however, we need `CharP.Lemmas` to be able to import some part of `ZMod`. -/ diff --git a/Mathlib/ModelTheory/Algebra/Field/CharP.lean b/Mathlib/ModelTheory/Algebra/Field/CharP.lean index 03491be723527..021b368387935 100644 --- a/Mathlib/ModelTheory/Algebra/Field/CharP.lean +++ b/Mathlib/ModelTheory/Algebra/Field/CharP.lean @@ -4,8 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes -/ -import Mathlib.Algebra.CharP.Defs -import Mathlib.Data.Nat.Prime.Defs +import Mathlib.Algebra.CharP.Basic import Mathlib.ModelTheory.Algebra.Ring.FreeCommRing import Mathlib.ModelTheory.Algebra.Field.Basic diff --git a/Mathlib/NumberTheory/MulChar/Basic.lean b/Mathlib/NumberTheory/MulChar/Basic.lean index 925ac23039f92..80a2bffa5b849 100644 --- a/Mathlib/NumberTheory/MulChar/Basic.lean +++ b/Mathlib/NumberTheory/MulChar/Basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Michael Stoll -/ import Mathlib.Algebra.CharP.Basic +import Mathlib.Algebra.CharP.Lemmas import Mathlib.Data.Fintype.Units import Mathlib.GroupTheory.OrderOfElement diff --git a/Mathlib/RingTheory/Polynomial/Dickson.lean b/Mathlib/RingTheory/Polynomial/Dickson.lean index 73b4540c4c50b..bf67e4ffe8272 100644 --- a/Mathlib/RingTheory/Polynomial/Dickson.lean +++ b/Mathlib/RingTheory/Polynomial/Dickson.lean @@ -7,7 +7,7 @@ import Mathlib.Algebra.CharP.Invertible import Mathlib.Data.ZMod.Basic import Mathlib.RingTheory.Localization.FractionRing import Mathlib.RingTheory.Polynomial.Chebyshev -import Mathlib.Algebra.CharP.Basic +import Mathlib.Algebra.CharP.Lemmas import Mathlib.Algebra.EuclideanDomain.Field import Mathlib.Algebra.Polynomial.Roots From 0b6c6c4d8f70155aae4dc62d24758734babe479a Mon Sep 17 00:00:00 2001 From: FR Date: Thu, 24 Oct 2024 21:25:43 +0000 Subject: [PATCH 395/505] =?UTF-8?q?feat:=20use=20aliases=20for=20`Covarian?= =?UTF-8?q?tClass=20=CE=B1=20=CE=B1=20(=C2=B7=20*=20=C2=B7)=20(=C2=B7=20?= =?UTF-8?q?=E2=89=A4=20=C2=B7)`=20etc=20(#13154)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- .../ZeroDivisorsInAddMonoidAlgebras.lean | 22 +- .../Algebra/BigOperators/Group/Finset.lean | 4 +- Mathlib/Algebra/DirectSum/Internal.lean | 2 +- .../Group/Pointwise/Finset/Interval.lean | 4 +- Mathlib/Algebra/Group/UniqueProds/Basic.lean | 6 +- Mathlib/Algebra/MonoidAlgebra/Degree.lean | 31 +- Mathlib/Algebra/Order/Antidiag/Prod.lean | 2 +- Mathlib/Algebra/Order/Archimedean/Basic.lean | 2 +- .../Order/BigOperators/Group/Finset.lean | 2 +- .../Order/BigOperators/Group/List.lean | 52 +-- Mathlib/Algebra/Order/Group/Abs.lean | 6 +- .../Algebra/Order/Group/CompleteLattice.lean | 4 +- Mathlib/Algebra/Order/Group/Defs.lean | 18 +- .../Algebra/Order/Group/DenselyOrdered.lean | 2 +- Mathlib/Algebra/Order/Group/Lattice.lean | 22 +- Mathlib/Algebra/Order/Group/MinMax.lean | 8 +- Mathlib/Algebra/Order/Group/OrderIso.lean | 7 +- .../Algebra/Order/Group/Pointwise/Bounds.lean | 8 +- .../Group/Pointwise/CompleteLattice.lean | 4 +- Mathlib/Algebra/Order/Group/PosPart.lean | 18 +- .../Algebra/Order/Group/Unbundled/Abs.lean | 10 +- .../Algebra/Order/Group/Unbundled/Basic.lean | 91 ++-- .../Order/GroupWithZero/Canonical.lean | 14 +- .../Algebra/Order/GroupWithZero/WithZero.lean | 8 +- Mathlib/Algebra/Order/Hom/Monoid.lean | 2 +- Mathlib/Algebra/Order/Interval/Basic.lean | 17 +- Mathlib/Algebra/Order/Kleene.lean | 6 +- Mathlib/Algebra/Order/Module/OrderedSMul.lean | 4 +- Mathlib/Algebra/Order/Monoid/Basic.lean | 4 +- .../Algebra/Order/Monoid/Canonical/Defs.lean | 2 +- Mathlib/Algebra/Order/Monoid/Defs.lean | 18 +- Mathlib/Algebra/Order/Monoid/NatCast.lean | 18 +- Mathlib/Algebra/Order/Monoid/OrderDual.lean | 6 +- Mathlib/Algebra/Order/Monoid/Prod.lean | 2 +- Mathlib/Algebra/Order/Monoid/Submonoid.lean | 2 +- .../Algebra/Order/Monoid/Unbundled/Basic.lean | 395 +++++++++--------- .../Algebra/Order/Monoid/Unbundled/Defs.lean | 24 +- .../Order/Monoid/Unbundled/ExistsOfLE.lean | 16 +- .../Order/Monoid/Unbundled/MinMax.lean | 31 +- .../Order/Monoid/Unbundled/OrderDual.lean | 28 +- .../Algebra/Order/Monoid/Unbundled/Pow.lean | 34 +- .../Order/Monoid/Unbundled/WithTop.lean | 96 ++--- Mathlib/Algebra/Order/Positive/Ring.lean | 31 +- Mathlib/Algebra/Order/Ring/Basic.lean | 2 +- Mathlib/Algebra/Order/Ring/Canonical.lean | 6 +- Mathlib/Algebra/Order/Ring/Cast.lean | 2 +- .../Algebra/Order/Ring/Unbundled/Basic.lean | 158 +++---- .../Algebra/Order/Ring/Unbundled/Nonneg.lean | 25 +- Mathlib/Algebra/Order/Ring/Unbundled/Rat.lean | 2 +- Mathlib/Algebra/Order/Star/Basic.lean | 4 +- Mathlib/Algebra/Order/Sub/Basic.lean | 7 +- Mathlib/Algebra/Order/Sub/Defs.lean | 22 +- .../Algebra/Order/Sub/Unbundled/Basic.lean | 12 +- Mathlib/Algebra/Order/Sub/Unbundled/Hom.lean | 4 +- Mathlib/Algebra/Tropical/Basic.lean | 26 +- .../SetFamily/CauchyDavenport.lean | 2 +- Mathlib/Data/DFinsupp/Lex.lean | 26 +- Mathlib/Data/DFinsupp/Order.lean | 4 +- Mathlib/Data/ENNReal/Basic.lean | 4 +- Mathlib/Data/ENNReal/Operations.lean | 8 +- Mathlib/Data/Finset/Fold.lean | 2 +- Mathlib/Data/Finset/Lattice.lean | 4 +- Mathlib/Data/Finsupp/Lex.lean | 26 +- Mathlib/Data/Finsupp/Order.lean | 4 +- Mathlib/Data/Multiset/Basic.lean | 4 +- Mathlib/Data/NNReal/Defs.lean | 6 +- Mathlib/Data/Nat/Cast/Order/Basic.lean | 2 +- Mathlib/Data/Nat/Cast/Order/Ring.lean | 4 +- Mathlib/Data/PNat/Basic.lean | 16 +- Mathlib/Data/Set/MulAntidiagonal.lean | 6 +- Mathlib/Data/Set/Pointwise/Interval.lean | 26 +- Mathlib/Data/Set/Semiring.lean | 10 +- Mathlib/Data/Sign.lean | 4 +- Mathlib/LinearAlgebra/Matrix/PosDef.lean | 8 +- .../LinearAlgebra/QuadraticForm/Basic.lean | 2 +- Mathlib/LinearAlgebra/QuadraticForm/Prod.lean | 6 +- Mathlib/MeasureTheory/Function/LpOrder.lean | 2 +- .../Function/SimpleFuncDenseLp.lean | 3 +- .../MeasureTheory/Measure/MeasureSpace.lean | 3 +- Mathlib/MeasureTheory/Measure/Restrict.lean | 3 +- .../MeasureTheory/Measure/VectorMeasure.lean | 5 +- .../ConditionallyCompleteLattice/Group.lean | 14 +- Mathlib/Order/Filter/Pointwise.lean | 4 +- Mathlib/Order/Filter/Ring.lean | 4 +- Mathlib/Order/LiminfLimsup.lean | 12 +- Mathlib/Probability/Martingale/Basic.lean | 20 +- Mathlib/Probability/Process/Stopping.lean | 4 +- Mathlib/RingTheory/GradedAlgebra/Basic.lean | 2 +- Mathlib/SetTheory/Cardinal/Basic.lean | 4 +- Mathlib/SetTheory/Game/Basic.lean | 10 +- Mathlib/SetTheory/Game/PGame.lean | 8 +- Mathlib/SetTheory/Ordinal/Arithmetic.lean | 16 +- Mathlib/SetTheory/Ordinal/Basic.lean | 5 +- Mathlib/SetTheory/Ordinal/NaturalOps.lean | 7 +- Mathlib/Tactic/Bound.lean | 2 +- Mathlib/Tactic/GCongr/Core.lean | 4 +- Mathlib/Tactic/Positivity/Basic.lean | 10 +- Mathlib/Tactic/Ring/Compare.lean | 2 +- .../Topology/Algebra/Order/LiminfLimsup.lean | 12 +- Mathlib/Topology/ContinuousMap/Algebra.lean | 10 +- Mathlib/Topology/Support.lean | 3 +- 101 files changed, 806 insertions(+), 890 deletions(-) diff --git a/Counterexamples/ZeroDivisorsInAddMonoidAlgebras.lean b/Counterexamples/ZeroDivisorsInAddMonoidAlgebras.lean index 703b8f0eded4e..b13cc110a5128 100644 --- a/Counterexamples/ZeroDivisorsInAddMonoidAlgebras.lean +++ b/Counterexamples/ZeroDivisorsInAddMonoidAlgebras.lean @@ -28,8 +28,8 @@ nontrivial ring `R` and the additive group `G` with a torsion element can be any Besides this example, we also address a comment in `Data.Finsupp.Lex` to the effect that the proof that addition is monotone on `α →₀ N` uses that it is *strictly* monotone on `N`. -The specific statement is about `Finsupp.Lex.covariantClass_lt_left` and its analogue -`Finsupp.Lex.covariantClass_le_right`. We do not need two separate counterexamples, since the +The specific statement is about `Finsupp.Lex.addLeftStrictMono` and its analogue +`Finsupp.Lex.addRightStrictMono`. We do not need two separate counterexamples, since the operation is commutative. The example is very simple. Let `F = {0, 1}` with order determined by `0 < 1` and absorbing @@ -142,9 +142,9 @@ elab "guard_decl" na:ident mod:ident : command => do let .some mdni := env.getModuleIdx? mdn | throwError "the module {mod} is not imported!" unless dcli = mdni do throwError "instance {na} is no longer in {mod}." -guard_decl Finsupp.Lex.covariantClass_le_left Mathlib.Data.Finsupp.Lex +guard_decl Finsupp.Lex.addLeftMono Mathlib.Data.Finsupp.Lex -guard_decl Finsupp.Lex.covariantClass_le_right Mathlib.Data.Finsupp.Lex +guard_decl Finsupp.Lex.addRightMono Mathlib.Data.Finsupp.Lex namespace F @@ -182,20 +182,20 @@ instance : AddCommMonoid F where add_comm := by boom nsmul := nsmulRec -/-- The `CovariantClass`es asserting monotonicity of addition hold for `F`. -/ -instance covariantClass_add_le : CovariantClass F F (· + ·) (· ≤ ·) := +/-- The `AddLeftMono`es asserting monotonicity of addition hold for `F`. -/ +instance addLeftMono : AddLeftMono F := ⟨by boom⟩ -example : CovariantClass F F (Function.swap (· + ·)) (· ≤ ·) := by infer_instance +example : AddRightMono F := by infer_instance /-- The following examples show that `F` has all the typeclasses used by -`Finsupp.Lex.covariantClass_le_left`... -/ +`Finsupp.Lex.addLeftStrictMono`... -/ example : LinearOrder F := by infer_instance example : AddMonoid F := by infer_instance /-- ... except for the strict monotonicity of addition, the crux of the matter. -/ -example : ¬CovariantClass F F (· + ·) (· < ·) := fun h => +example : ¬AddLeftStrictMono F := fun h => lt_irrefl 1 <| (h.elim : Covariant F F (· + ·) (· < ·)) 1 z01 /-- A few `simp`-lemmas to take care of trivialities in the proof of the example below. -/ @@ -220,8 +220,8 @@ theorem f110 : ofLex (Finsupp.single (1 : F) (1 : F)) 0 = 0 := /-- Here we see that (not-necessarily strict) monotonicity of addition on `Lex (F →₀ F)` is not a consequence of monotonicity of addition on `F`. Strict monotonicity of addition on `F` is -enough and is the content of `Finsupp.Lex.covariantClass_le_left`. -/ -example : ¬CovariantClass (Lex (F →₀ F)) (Lex (F →₀ F)) (· + ·) (· ≤ ·) := by +enough and is the content of `Finsupp.Lex.addLeftStrictMono`. -/ +example : ¬AddLeftMono (Lex (F →₀ F)) := by rintro ⟨h⟩ refine (not_lt (α := Lex (F →₀ F))).mpr (@h (Finsupp.single (0 : F) (1 : F)) (Finsupp.single 1 1) (Finsupp.single 0 1) ?_) ⟨1, ?_⟩ diff --git a/Mathlib/Algebra/BigOperators/Group/Finset.lean b/Mathlib/Algebra/BigOperators/Group/Finset.lean index 4caade7753513..a15dd8845f035 100644 --- a/Mathlib/Algebra/BigOperators/Group/Finset.lean +++ b/Mathlib/Algebra/BigOperators/Group/Finset.lean @@ -1466,7 +1466,7 @@ reduces to the difference of the last and first terms when the function we are summing is monotone. -/ theorem sum_range_tsub [AddCommMonoid α] [PartialOrder α] [Sub α] [OrderedSub α] - [CovariantClass α α (· + ·) (· ≤ ·)] [ContravariantClass α α (· + ·) (· ≤ ·)] [ExistsAddOfLE α] + [AddLeftMono α] [AddLeftReflectLE α] [ExistsAddOfLE α] {f : ℕ → α} (h : Monotone f) (n : ℕ) : ∑ i ∈ range n, (f (i + 1) - f i) = f n - f 0 := by apply sum_range_induction @@ -1701,7 +1701,7 @@ theorem prod_ite_one (s : Finset α) (p : α → Prop) [DecidablePred p] @[to_additive] theorem prod_erase_lt_of_one_lt {γ : Type*} [DecidableEq α] [CommMonoid γ] [Preorder γ] - [CovariantClass γ γ (· * ·) (· < ·)] {s : Finset α} {d : α} (hd : d ∈ s) {f : α → γ} + [MulLeftStrictMono γ] {s : Finset α} {d : α} (hd : d ∈ s) {f : α → γ} (hdf : 1 < f d) : ∏ m ∈ s.erase d, f m < ∏ m ∈ s, f m := by conv in ∏ m ∈ s, f m => rw [← Finset.insert_erase hd] rw [Finset.prod_insert (Finset.not_mem_erase d s)] diff --git a/Mathlib/Algebra/DirectSum/Internal.lean b/Mathlib/Algebra/DirectSum/Internal.lean index 99122f4b761d9..aaf84929ec1cc 100644 --- a/Mathlib/Algebra/DirectSum/Internal.lean +++ b/Mathlib/Algebra/DirectSum/Internal.lean @@ -237,7 +237,7 @@ theorem coe_mul_of_apply_of_not_le (r : ⨁ i, A i) {i : ι} (r' : A i) (n : ι) · rw [DFinsupp.sum, Finset.sum_ite_of_false, Finset.sum_const_zero] exact fun x _ H => h ((self_le_add_left i x).trans_eq H) -variable [Sub ι] [OrderedSub ι] [ContravariantClass ι ι (· + ·) (· ≤ ·)] +variable [Sub ι] [OrderedSub ι] [AddLeftReflectLE ι] /- The following two lemmas only require the same hypotheses as `eq_tsub_iff_add_eq_of_le`, but we state them for `CanonicallyOrderedAddCommMonoid` + the above three typeclasses for convenience. -/ diff --git a/Mathlib/Algebra/Group/Pointwise/Finset/Interval.lean b/Mathlib/Algebra/Group/Pointwise/Finset/Interval.lean index fd9cc57c0dd5b..479121a3d4a1b 100644 --- a/Mathlib/Algebra/Group/Pointwise/Finset/Interval.lean +++ b/Mathlib/Algebra/Group/Pointwise/Finset/Interval.lean @@ -31,7 +31,7 @@ the unprimed names have been reserved for section ContravariantLE variable [Mul α] [Preorder α] [DecidableEq α] -variable [CovariantClass α α (· * ·) (· ≤ ·)] [CovariantClass α α (Function.swap HMul.hMul) LE.le] +variable [MulLeftMono α] [MulRightMono α] @[to_additive Icc_add_Icc_subset] theorem Icc_mul_Icc_subset' [LocallyFiniteOrder α] (a b c d : α) : @@ -51,7 +51,7 @@ end ContravariantLE section ContravariantLT variable [Mul α] [PartialOrder α] [DecidableEq α] -variable [CovariantClass α α (· * ·) (· < ·)] [CovariantClass α α (Function.swap HMul.hMul) LT.lt] +variable [MulLeftStrictMono α] [MulRightStrictMono α] @[to_additive Icc_add_Ico_subset] theorem Icc_mul_Ico_subset' [LocallyFiniteOrder α] (a b c d : α) : diff --git a/Mathlib/Algebra/Group/UniqueProds/Basic.lean b/Mathlib/Algebra/Group/UniqueProds/Basic.lean index f43ae9290e9cb..78c9df91994a4 100644 --- a/Mathlib/Algebra/Group/UniqueProds/Basic.lean +++ b/Mathlib/Algebra/Group/UniqueProds/Basic.lean @@ -581,7 +581,7 @@ theorem of_mulOpposite (h : TwoUniqueProds Gᵐᵒᵖ) : TwoUniqueProds G where "This instance asserts that if `G` has a right-cancellative addition, a linear order, and addition is strictly monotone w.r.t. the second argument, then `G` has `TwoUniqueSums`." ] instance (priority := 100) of_covariant_right [IsRightCancelMul G] - [LinearOrder G] [CovariantClass G G (· * ·) (· < ·)] : + [LinearOrder G] [MulLeftStrictMono G] : TwoUniqueProds G where uniqueMul_of_one_lt_card {A B} hc := by obtain ⟨hA, hB, -⟩ := Nat.one_lt_mul_iff.mp hc @@ -615,10 +615,10 @@ open MulOpposite in "This instance asserts that if `G` has a left-cancellative addition, a linear order, and addition is strictly monotone w.r.t. the first argument, then `G` has `TwoUniqueSums`." ] instance (priority := 100) of_covariant_left [IsLeftCancelMul G] - [LinearOrder G] [CovariantClass G G (Function.swap (· * ·)) (· < ·)] : + [LinearOrder G] [MulRightStrictMono G] : TwoUniqueProds G := let _ := LinearOrder.lift' (unop : Gᵐᵒᵖ → G) unop_injective - let _ : CovariantClass Gᵐᵒᵖ Gᵐᵒᵖ (· * ·) (· < ·) := + let _ : MulLeftStrictMono Gᵐᵒᵖ := { elim := fun _ _ _ bc ↦ mul_lt_mul_right' (α := G) bc (unop _) } of_mulOpposite of_covariant_right diff --git a/Mathlib/Algebra/MonoidAlgebra/Degree.lean b/Mathlib/Algebra/MonoidAlgebra/Degree.lean index 4546eb21d1d13..616114af6f493 100644 --- a/Mathlib/Algebra/MonoidAlgebra/Degree.lean +++ b/Mathlib/Algebra/MonoidAlgebra/Degree.lean @@ -93,9 +93,8 @@ end ExplicitDegrees section AddOnly -variable [Add A] [Add B] [Add T] [CovariantClass B B (· + ·) (· ≤ ·)] - [CovariantClass B B (Function.swap (· + ·)) (· ≤ ·)] [CovariantClass T T (· + ·) (· ≤ ·)] - [CovariantClass T T (Function.swap (· + ·)) (· ≤ ·)] +variable [Add A] [Add B] [Add T] [AddLeftMono B] [AddRightMono B] + [AddLeftMono T] [AddRightMono T] theorem sup_support_mul_le {degb : A → B} (degbm : ∀ {a b}, degb (a + b) ≤ degb a + degb b) (f g : R[A]) : @@ -113,9 +112,8 @@ end AddOnly section AddMonoids -variable [AddMonoid A] [AddMonoid B] [CovariantClass B B (· + ·) (· ≤ ·)] - [CovariantClass B B (Function.swap (· + ·)) (· ≤ ·)] [AddMonoid T] - [CovariantClass T T (· + ·) (· ≤ ·)] [CovariantClass T T (Function.swap (· + ·)) (· ≤ ·)] +variable [AddMonoid A] [AddMonoid B] [AddLeftMono B] [AddRightMono B] + [AddMonoid T] [AddLeftMono T] [AddRightMono T] {degb : A → B} {degt : A → T} theorem sup_support_list_prod_le (degb0 : degb 0 ≤ 0) @@ -159,9 +157,8 @@ end Semiring section CommutativeLemmas -variable [CommSemiring R] [AddCommMonoid A] [AddCommMonoid B] [CovariantClass B B (· + ·) (· ≤ ·)] - [CovariantClass B B (Function.swap (· + ·)) (· ≤ ·)] [AddCommMonoid T] - [CovariantClass T T (· + ·) (· ≤ ·)] [CovariantClass T T (Function.swap (· + ·)) (· ≤ ·)] +variable [CommSemiring R] [AddCommMonoid A] [AddCommMonoid B] [AddLeftMono B] [AddRightMono B] + [AddCommMonoid T] [AddLeftMono T] [AddRightMono T] {degb : A → B} {degt : A → T} theorem sup_support_multiset_prod_le (degb0 : degb 0 ≤ 0) @@ -285,13 +282,13 @@ theorem supDegree_eq_of_max {b : B} (hb : b ∈ Set.range D) (hmem : D.invFun b variable [Add B] theorem supDegree_mul_le (hadd : ∀ a1 a2, D (a1 + a2) = D a1 + D a2) - [CovariantClass B B (· + ·) (· ≤ ·)] [CovariantClass B B (Function.swap (· + ·)) (· ≤ ·)] : + [AddLeftMono B] [AddRightMono B] : (p * q).supDegree D ≤ p.supDegree D + q.supDegree D := sup_support_mul_le (fun {_ _} => (hadd _ _).le) p q theorem supDegree_prod_le {R A B : Type*} [CommSemiring R] [AddCommMonoid A] [AddCommMonoid B] [SemilatticeSup B] [OrderBot B] - [CovariantClass B B (· + ·) (· ≤ ·)] [CovariantClass B B (Function.swap (· + ·)) (· ≤ ·)] + [AddLeftMono B] [AddRightMono B] {D : A → B} (hzero : D 0 = 0) (hadd : ∀ a1 a2, D (a1 + a2) = D a1 + D a2) {ι} {s : Finset ι} {f : ι → R[A]} : (∏ i ∈ s, f i).supDegree D ≤ ∑ i ∈ s, (f i).supDegree D := by @@ -304,7 +301,7 @@ theorem supDegree_prod_le {R A B : Type*} [CommSemiring R] [AddCommMonoid A] [Ad exact (supDegree_mul_le hadd).trans (add_le_add_left ih _) theorem apply_add_of_supDegree_le (hadd : ∀ a1 a2, D (a1 + a2) = D a1 + D a2) - [CovariantClass B B (· + ·) (· < ·)] [CovariantClass B B (Function.swap (· + ·)) (· < ·)] + [AddLeftStrictMono B] [AddRightStrictMono B] (hD : D.Injective) {ap aq : A} (hp : p.supDegree D ≤ D ap) (hq : q.supDegree D ≤ D aq) : (p * q) (ap + aq) = p ap * q aq := by classical @@ -317,7 +314,7 @@ theorem apply_add_of_supDegree_le (hadd : ∀ a1 a2, D (a1 + a2) = D a1 + D a2) · refine fun a ha hne => Finset.sum_eq_zero (fun a' ha' => if_neg <| fun he => ?_) apply_fun D at he simp_rw [hadd] at he - have := covariantClass_le_of_lt B B (· + ·) + have := addLeftMono_of_addLeftStrictMono B exact (add_lt_add_of_lt_of_le (((Finset.le_sup ha).trans hp).lt_of_ne <| hD.ne_iff.2 hne) <| (Finset.le_sup ha').trans hq).ne he · refine fun h => Finset.sum_eq_zero (fun a _ => ite_eq_right_iff.mpr <| fun _ => ?_) @@ -459,7 +456,7 @@ lemma sum_ne_zero_of_injOn_supDegree (hs : s ≠ ∅) sum_ne_zero_of_injOn_supDegree' ⟨i, hi, hf i hi⟩ hd variable [Add B] -variable [CovariantClass B B (· + ·) (· < ·)] [CovariantClass B B (Function.swap (· + ·)) (· < ·)] +variable [AddLeftStrictMono B] [AddRightStrictMono B] lemma apply_supDegree_add_supDegree (hD : D.Injective) (hadd : ∀ a1 a2, D (a1 + a2) = D a1 + D a2) : (p * q) (D.invFun (p.supDegree D + q.supDegree D)) = p.leadingCoeff D * q.leadingCoeff D := by @@ -545,7 +542,7 @@ lemma Monic.mul section AddMonoid variable {A B : Type*} [AddMonoid A] [AddMonoid B] [LinearOrder B] [OrderBot B] - [CovariantClass B B (· + ·) (· < ·)] [CovariantClass B B (Function.swap (· + ·)) (· < ·)] + [AddLeftStrictMono B] [AddRightStrictMono B] {D : A → B} {p : R[A]} {n : ℕ} lemma Monic.pow @@ -592,9 +589,7 @@ theorem infDegree_withTop_some_comp {s : AddMonoidAlgebra R A} (hs : s.support.N unfold AddMonoidAlgebra.infDegree rw [← Finset.coe_inf' hs, Finset.inf'_eq_inf] -theorem le_infDegree_mul - [AddZeroClass A] [Add T] - [CovariantClass T T (· + ·) (· ≤ ·)] [CovariantClass T T (Function.swap (· + ·)) (· ≤ ·)] +theorem le_infDegree_mul [AddZeroClass A] [Add T] [AddLeftMono T] [AddRightMono T] (D : AddHom A T) (f g : R[A]) : f.infDegree D + g.infDegree D ≤ (f * g).infDegree D := le_inf_support_mul (fun {a b : A} => (map_add D a b).ge) _ _ diff --git a/Mathlib/Algebra/Order/Antidiag/Prod.lean b/Mathlib/Algebra/Order/Antidiag/Prod.lean index b41df751f3663..1bd46120668a9 100644 --- a/Mathlib/Algebra/Order/Antidiag/Prod.lean +++ b/Mathlib/Algebra/Order/Antidiag/Prod.lean @@ -144,7 +144,7 @@ end CanonicallyOrderedAddCommMonoid section OrderedSub variable [CanonicallyOrderedAddCommMonoid A] [Sub A] [OrderedSub A] -variable [ContravariantClass A A (· + ·) (· ≤ ·)] +variable [AddLeftReflectLE A] variable [HasAntidiagonal A] theorem filter_fst_eq_antidiagonal (n m : A) [DecidablePred (· = m)] [Decidable (m ≤ n)] : diff --git a/Mathlib/Algebra/Order/Archimedean/Basic.lean b/Mathlib/Algebra/Order/Archimedean/Basic.lean index 4d20bc04708f5..fdb47f471b9bb 100644 --- a/Mathlib/Algebra/Order/Archimedean/Basic.lean +++ b/Mathlib/Algebra/Order/Archimedean/Basic.lean @@ -71,7 +71,7 @@ variable {M : Type*} @[to_additive] theorem exists_lt_pow [OrderedCommMonoid M] [MulArchimedean M] - [CovariantClass M M (· * ·) (· < ·)] {a : M} (ha : 1 < a) (b : M) : + [MulLeftStrictMono M] {a : M} (ha : 1 < a) (b : M) : ∃ n : ℕ, b < a ^ n := let ⟨k, hk⟩ := MulArchimedean.arch b ha ⟨k + 1, hk.trans_lt <| pow_lt_pow_right' ha k.lt_succ_self⟩ diff --git a/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean b/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean index c0119e6787e06..67e509080c462 100644 --- a/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean +++ b/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean @@ -231,7 +231,7 @@ theorem abs_sum_of_nonneg' {G : Type*} [LinearOrderedAddCommGroup G] {f : ι → rw [abs_of_nonneg (Finset.sum_nonneg' hf)] section CommMonoid -variable [CommMonoid α] [LE α] [CovariantClass α α (· * ·) (· ≤ ·)] {s : Finset ι} {f : ι → α} +variable [CommMonoid α] [LE α] [MulLeftMono α] {s : Finset ι} {f : ι → α} @[to_additive (attr := simp)] lemma mulLECancellable_prod : diff --git a/Mathlib/Algebra/Order/BigOperators/Group/List.lean b/Mathlib/Algebra/Order/BigOperators/Group/List.lean index 94dd5a115870c..bf44e7ef2b251 100644 --- a/Mathlib/Algebra/Order/BigOperators/Group/List.lean +++ b/Mathlib/Algebra/Order/BigOperators/Group/List.lean @@ -21,8 +21,8 @@ section Monoid variable [Monoid M] @[to_additive sum_le_sum] -lemma Forall₂.prod_le_prod' [Preorder M] [CovariantClass M M (Function.swap (· * ·)) (· ≤ ·)] - [CovariantClass M M (· * ·) (· ≤ ·)] {l₁ l₂ : List M} (h : Forall₂ (· ≤ ·) l₁ l₂) : +lemma Forall₂.prod_le_prod' [Preorder M] [MulRightMono M] + [MulLeftMono M] {l₁ l₂ : List M} (h : Forall₂ (· ≤ ·) l₁ l₂) : l₁.prod ≤ l₂.prod := by induction' h with a b la lb hab ih ih' · rfl @@ -35,8 +35,8 @@ of `∀ a ∈ l₂, 1 ≤ a` but this lemma is not yet in `mathlib`. -/ then `l₁.sum ≤ l₂.sum`. One can prove a stronger version assuming `∀ a ∈ l₂.diff l₁, 0 ≤ a` instead of `∀ a ∈ l₂, 0 ≤ a` but this lemma is not yet in `mathlib`."] -lemma Sublist.prod_le_prod' [Preorder M] [CovariantClass M M (Function.swap (· * ·)) (· ≤ ·)] - [CovariantClass M M (· * ·) (· ≤ ·)] {l₁ l₂ : List M} (h : l₁ <+ l₂) +lemma Sublist.prod_le_prod' [Preorder M] [MulRightMono M] + [MulLeftMono M] {l₁ l₂ : List M} (h : l₁ <+ l₂) (h₁ : ∀ a ∈ l₂, (1 : M) ≤ a) : l₁.prod ≤ l₂.prod := by induction h with | slnil => rfl @@ -49,22 +49,22 @@ lemma Sublist.prod_le_prod' [Preorder M] [CovariantClass M M (Function.swap (· @[to_additive sum_le_sum] lemma SublistForall₂.prod_le_prod' [Preorder M] - [CovariantClass M M (Function.swap (· * ·)) (· ≤ ·)] [CovariantClass M M (· * ·) (· ≤ ·)] + [MulRightMono M] [MulLeftMono M] {l₁ l₂ : List M} (h : SublistForall₂ (· ≤ ·) l₁ l₂) (h₁ : ∀ a ∈ l₂, (1 : M) ≤ a) : l₁.prod ≤ l₂.prod := let ⟨_, hall, hsub⟩ := sublistForall₂_iff.1 h hall.prod_le_prod'.trans <| hsub.prod_le_prod' h₁ @[to_additive sum_le_sum] -lemma prod_le_prod' [Preorder M] [CovariantClass M M (Function.swap (· * ·)) (· ≤ ·)] - [CovariantClass M M (· * ·) (· ≤ ·)] {l : List ι} {f g : ι → M} (h : ∀ i ∈ l, f i ≤ g i) : +lemma prod_le_prod' [Preorder M] [MulRightMono M] + [MulLeftMono M] {l : List ι} {f g : ι → M} (h : ∀ i ∈ l, f i ≤ g i) : (l.map f).prod ≤ (l.map g).prod := Forall₂.prod_le_prod' <| by simpa @[to_additive sum_lt_sum] -lemma prod_lt_prod' [Preorder M] [CovariantClass M M (· * ·) (· < ·)] - [CovariantClass M M (· * ·) (· ≤ ·)] [CovariantClass M M (Function.swap (· * ·)) (· < ·)] - [CovariantClass M M (Function.swap (· * ·)) (· ≤ ·)] {l : List ι} (f g : ι → M) +lemma prod_lt_prod' [Preorder M] [MulLeftStrictMono M] + [MulLeftMono M] [MulRightStrictMono M] + [MulRightMono M] {l : List ι} (f g : ι → M) (h₁ : ∀ i ∈ l, f i ≤ g i) (h₂ : ∃ i ∈ l, f i < g i) : (l.map f).prod < (l.map g).prod := by induction' l with i l ihl · rcases h₂ with ⟨_, ⟨⟩, _⟩ @@ -75,42 +75,42 @@ lemma prod_lt_prod' [Preorder M] [CovariantClass M M (· * ·) (· < ·)] · exact mul_lt_mul_of_le_of_lt h₁.1 <| ihl h₁.2 ‹_› @[to_additive] -lemma prod_lt_prod_of_ne_nil [Preorder M] [CovariantClass M M (· * ·) (· < ·)] - [CovariantClass M M (· * ·) (· ≤ ·)] [CovariantClass M M (Function.swap (· * ·)) (· < ·)] - [CovariantClass M M (Function.swap (· * ·)) (· ≤ ·)] {l : List ι} (hl : l ≠ []) (f g : ι → M) +lemma prod_lt_prod_of_ne_nil [Preorder M] [MulLeftStrictMono M] + [MulLeftMono M] [MulRightStrictMono M] + [MulRightMono M] {l : List ι} (hl : l ≠ []) (f g : ι → M) (hlt : ∀ i ∈ l, f i < g i) : (l.map f).prod < (l.map g).prod := (prod_lt_prod' f g fun i hi => (hlt i hi).le) <| (exists_mem_of_ne_nil l hl).imp fun i hi => ⟨hi, hlt i hi⟩ @[to_additive sum_le_card_nsmul] -lemma prod_le_pow_card [Preorder M] [CovariantClass M M (Function.swap (· * ·)) (· ≤ ·)] - [CovariantClass M M (· * ·) (· ≤ ·)] (l : List M) (n : M) (h : ∀ x ∈ l, x ≤ n) : +lemma prod_le_pow_card [Preorder M] [MulRightMono M] + [MulLeftMono M] (l : List M) (n : M) (h : ∀ x ∈ l, x ≤ n) : l.prod ≤ n ^ l.length := by simpa only [map_id', map_const', prod_replicate] using prod_le_prod' h @[to_additive card_nsmul_le_sum] -lemma pow_card_le_prod [Preorder M] [CovariantClass M M (Function.swap (· * ·)) (· ≤ ·)] - [CovariantClass M M (· * ·) (· ≤ ·)] (l : List M) (n : M) (h : ∀ x ∈ l, n ≤ x) : +lemma pow_card_le_prod [Preorder M] [MulRightMono M] + [MulLeftMono M] (l : List M) (n : M) (h : ∀ x ∈ l, n ≤ x) : n ^ l.length ≤ l.prod := @prod_le_pow_card Mᵒᵈ _ _ _ _ l n h @[to_additive exists_lt_of_sum_lt] -lemma exists_lt_of_prod_lt' [LinearOrder M] [CovariantClass M M (Function.swap (· * ·)) (· ≤ ·)] - [CovariantClass M M (· * ·) (· ≤ ·)] {l : List ι} (f g : ι → M) +lemma exists_lt_of_prod_lt' [LinearOrder M] [MulRightMono M] + [MulLeftMono M] {l : List ι} (f g : ι → M) (h : (l.map f).prod < (l.map g).prod) : ∃ i ∈ l, f i < g i := by contrapose! h exact prod_le_prod' h @[to_additive exists_le_of_sum_le] -lemma exists_le_of_prod_le' [LinearOrder M] [CovariantClass M M (· * ·) (· < ·)] - [CovariantClass M M (· * ·) (· ≤ ·)] [CovariantClass M M (Function.swap (· * ·)) (· < ·)] - [CovariantClass M M (Function.swap (· * ·)) (· ≤ ·)] {l : List ι} (hl : l ≠ []) (f g : ι → M) +lemma exists_le_of_prod_le' [LinearOrder M] [MulLeftStrictMono M] + [MulLeftMono M] [MulRightStrictMono M] + [MulRightMono M] {l : List ι} (hl : l ≠ []) (f g : ι → M) (h : (l.map f).prod ≤ (l.map g).prod) : ∃ x ∈ l, f x ≤ g x := by contrapose! h exact prod_lt_prod_of_ne_nil hl _ _ h @[to_additive sum_nonneg] -lemma one_le_prod_of_one_le [Preorder M] [CovariantClass M M (· * ·) (· ≤ ·)] {l : List M} +lemma one_le_prod_of_one_le [Preorder M] [MulLeftMono M] {l : List M} (hl₁ : ∀ x ∈ l, (1 : M) ≤ x) : 1 ≤ l.prod := by -- We don't use `pow_card_le_prod` to avoid assumption -- [covariant_class M M (function.swap (*)) (≤)] @@ -121,7 +121,7 @@ lemma one_le_prod_of_one_le [Preorder M] [CovariantClass M M (· * ·) (· ≤ @[to_additive] lemma max_prod_le (l : List α) (f g : α → M) [LinearOrder M] - [CovariantClass M M (· * ·) (· ≤ ·)] [CovariantClass M M (Function.swap (· * ·)) (· ≤ ·)] : + [MulLeftMono M] [MulRightMono M] : 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 @@ -129,8 +129,8 @@ lemma max_prod_le (l : List α) (f g : α → M) [LinearOrder M] · 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) : +lemma prod_min_le [LinearOrder M] [MulLeftMono M] + [MulRightMono M] (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 diff --git a/Mathlib/Algebra/Order/Group/Abs.lean b/Mathlib/Algebra/Order/Group/Abs.lean index 12f0c643f7f22..06283c5ee633b 100644 --- a/Mathlib/Algebra/Order/Group/Abs.lean +++ b/Mathlib/Algebra/Order/Group/Abs.lean @@ -65,7 +65,7 @@ variable [LinearOrderedAddCommGroup α] {a b c : α} -- Porting note: -- Lean can perfectly well find this instance, -- but in the rewrites below it is going looking for it without having fixed `α`. -example : CovariantClass α α (swap fun x y ↦ x + y) fun x y ↦ x ≤ y := inferInstance +example : AddRightMono α := inferInstance theorem abs_le : |a| ≤ b ↔ -b ≤ a ∧ a ≤ b := by rw [abs_le', and_comm, @neg_le α] @@ -79,14 +79,14 @@ theorem le_of_abs_le (h : |a| ≤ b) : a ≤ b := @[to_additive] theorem apply_abs_le_mul_of_one_le' {β : Type*} [MulOneClass β] [Preorder β] - [CovariantClass β β (· * ·) (· ≤ ·)] [CovariantClass β β (swap (· * ·)) (· ≤ ·)] {f : α → β} + [MulLeftMono β] [MulRightMono β] {f : α → β} {a : α} (h₁ : 1 ≤ f a) (h₂ : 1 ≤ f (-a)) : f |a| ≤ f a * f (-a) := (le_total a 0).rec (fun ha => (abs_of_nonpos ha).symm ▸ le_mul_of_one_le_left' h₁) fun ha => (abs_of_nonneg ha).symm ▸ le_mul_of_one_le_right' h₂ @[to_additive] theorem apply_abs_le_mul_of_one_le {β : Type*} [MulOneClass β] [Preorder β] - [CovariantClass β β (· * ·) (· ≤ ·)] [CovariantClass β β (swap (· * ·)) (· ≤ ·)] {f : α → β} + [MulLeftMono β] [MulRightMono β] {f : α → β} (h : ∀ x, 1 ≤ f x) (a : α) : f |a| ≤ f a * f (-a) := apply_abs_le_mul_of_one_le' (h _) (h _) diff --git a/Mathlib/Algebra/Order/Group/CompleteLattice.lean b/Mathlib/Algebra/Order/Group/CompleteLattice.lean index 78a88b90761f3..69e9649b537de 100644 --- a/Mathlib/Algebra/Order/Group/CompleteLattice.lean +++ b/Mathlib/Algebra/Order/Group/CompleteLattice.lean @@ -15,7 +15,7 @@ open Function Set variable {ι G : Type*} [Group G] [ConditionallyCompleteLattice G] [Nonempty ι] {f : ι → G} section Right -variable [CovariantClass G G (swap (· * ·)) (· ≤ ·)] +variable [MulRightMono G] @[to_additive] lemma ciSup_mul (hf : BddAbove (range f)) (a : G) : (⨆ i, f i) * a = ⨆ i, f i * a := @@ -36,7 +36,7 @@ lemma ciInf_div (hf : BddBelow (range f)) (a : G) : (⨅ i, f i) / a = ⨅ i, f end Right section Left -variable [CovariantClass G G (· * ·) (· ≤ ·)] +variable [MulLeftMono G] @[to_additive] lemma mul_ciSup (hf : BddAbove (range f)) (a : G) : (a * ⨆ i, f i) = ⨆ i, a * f i := diff --git a/Mathlib/Algebra/Order/Group/Defs.lean b/Mathlib/Algebra/Order/Group/Defs.lean index a75ae0afc6711..d18bd57fc7cfe 100644 --- a/Mathlib/Algebra/Order/Group/Defs.lean +++ b/Mathlib/Algebra/Order/Group/Defs.lean @@ -46,8 +46,8 @@ class OrderedCommGroup (α : Type u) extends CommGroup α, PartialOrder α where attribute [to_additive] OrderedCommGroup @[to_additive] -instance OrderedCommGroup.to_covariantClass_left_le (α : Type u) [OrderedCommGroup α] : - CovariantClass α α (· * ·) (· ≤ ·) where +instance OrderedCommGroup.toMulLeftMono (α : Type u) [OrderedCommGroup α] : + MulLeftMono α where elim a b c bc := OrderedCommGroup.mul_le_mul_left b c bc a -- See note [lower instance priority] @@ -56,8 +56,8 @@ instance (priority := 100) OrderedCommGroup.toOrderedCancelCommMonoid [OrderedCo OrderedCancelCommMonoid α := { ‹OrderedCommGroup α› with le_of_mul_le_mul_left := fun _ _ _ ↦ le_of_mul_le_mul_left' } -example (α : Type u) [OrderedAddCommGroup α] : CovariantClass α α (swap (· + ·)) (· < ·) := - IsRightCancelAdd.covariant_swap_add_lt_of_covariant_swap_add_le α +example (α : Type u) [OrderedAddCommGroup α] : AddRightStrictMono α := + inferInstance -- Porting note: this instance is not used, -- and causes timeouts after lean4#2210. @@ -65,17 +65,17 @@ example (α : Type u) [OrderedAddCommGroup α] : CovariantClass α α (swap (· -- but without the motivation clearly explained. /-- A choice-free shortcut instance. -/ @[to_additive "A choice-free shortcut instance."] -theorem OrderedCommGroup.to_contravariantClass_left_le (α : Type u) [OrderedCommGroup α] : - ContravariantClass α α (· * ·) (· ≤ ·) where +theorem OrderedCommGroup.toMulLeftReflectLE (α : Type u) [OrderedCommGroup α] : + MulLeftReflectLE α where elim a b c bc := by simpa using mul_le_mul_left' bc a⁻¹ -- Porting note: this instance is not used, -- and causes timeouts after lean4#2210. --- See further explanation on `OrderedCommGroup.to_contravariantClass_left_le`. +-- See further explanation on `OrderedCommGroup.toMulLeftReflectLE`. /-- A choice-free shortcut instance. -/ @[to_additive "A choice-free shortcut instance."] -theorem OrderedCommGroup.to_contravariantClass_right_le (α : Type u) [OrderedCommGroup α] : - ContravariantClass α α (swap (· * ·)) (· ≤ ·) where +theorem OrderedCommGroup.toMulRightReflectLE (α : Type u) [OrderedCommGroup α] : + MulRightReflectLE α where elim a b c bc := by simpa using mul_le_mul_right' bc a⁻¹ alias OrderedCommGroup.mul_lt_mul_left' := mul_lt_mul_left' diff --git a/Mathlib/Algebra/Order/Group/DenselyOrdered.lean b/Mathlib/Algebra/Order/Group/DenselyOrdered.lean index 2e25d60a9f7e1..70408fea662ab 100644 --- a/Mathlib/Algebra/Order/Group/DenselyOrdered.lean +++ b/Mathlib/Algebra/Order/Group/DenselyOrdered.lean @@ -17,7 +17,7 @@ variable {α : Type*} section DenselyOrdered variable [Group α] [LinearOrder α] -variable [CovariantClass α α (· * ·) (· ≤ ·)] +variable [MulLeftMono α] variable [DenselyOrdered α] {a b : α} @[to_additive] diff --git a/Mathlib/Algebra/Order/Group/Lattice.lean b/Mathlib/Algebra/Order/Group/Lattice.lean index af14c9db3c4d1..8260bbb4f5c92 100644 --- a/Mathlib/Algebra/Order/Group/Lattice.lean +++ b/Mathlib/Algebra/Order/Group/Lattice.lean @@ -14,8 +14,8 @@ underpinnings of vector lattices, Banach lattices, AL-space, AM-space etc. A lattice ordered group is a type `α` satisfying: * `Lattice α` * `CommGroup α` -* `CovariantClass α α (· * ·) (· ≤ ·)` -* `CovariantClass α α (swap (· * ·)) (· ≤ ·)` +* `MulLeftMono α` +* `MulRightMono α` This file establishes basic properties of lattice ordered groups. It is shown that when the group is commutative, the lattice is distributive. This also holds in the non-commutative case @@ -44,37 +44,37 @@ variable [Lattice α] [Group α] -- Special case of Bourbaki A.VI.9 (1) @[to_additive] -lemma mul_sup [CovariantClass α α (· * ·) (· ≤ ·)] (a b c : α) : +lemma mul_sup [MulLeftMono α] (a b c : α) : c * (a ⊔ b) = c * a ⊔ c * b := (OrderIso.mulLeft _).map_sup _ _ @[to_additive] -lemma sup_mul [CovariantClass α α (swap (· * ·)) (· ≤ ·)] (a b c : α) : +lemma sup_mul [MulRightMono α] (a b c : α) : (a ⊔ b) * c = a * c ⊔ b * c := (OrderIso.mulRight _).map_sup _ _ @[to_additive] -lemma mul_inf [CovariantClass α α (· * ·) (· ≤ ·)] (a b c : α) : +lemma mul_inf [MulLeftMono α] (a b c : α) : c * (a ⊓ b) = c * a ⊓ c * b := (OrderIso.mulLeft _).map_inf _ _ @[to_additive] -lemma inf_mul [CovariantClass α α (swap (· * ·)) (· ≤ ·)] (a b c : α) : +lemma inf_mul [MulRightMono α] (a b c : α) : (a ⊓ b) * c = a * c ⊓ b * c := (OrderIso.mulRight _).map_inf _ _ @[to_additive] -lemma sup_div [CovariantClass α α (swap (· * ·)) (· ≤ ·)] (a b c : α) : +lemma sup_div [MulRightMono α] (a b c : α) : (a ⊔ b) / c = a / c ⊔ b / c := (OrderIso.divRight _).map_sup _ _ @[to_additive] -lemma inf_div [CovariantClass α α (swap (· * ·)) (· ≤ ·)] (a b c : α) : +lemma inf_div [MulRightMono α] (a b c : α) : (a ⊓ b) / c = a / c ⊓ b / c := (OrderIso.divRight _).map_inf _ _ section -variable [CovariantClass α α (· * ·) (· ≤ ·)] [CovariantClass α α (swap (· * ·)) (· ≤ ·)] +variable [MulLeftMono α] [MulRightMono α] @[to_additive] lemma inv_sup (a b : α) : (a ⊔ b)⁻¹ = a⁻¹ ⊓ b⁻¹ := (OrderIso.inv α).map_sup _ _ @@ -106,7 +106,7 @@ variable [Lattice α] [CommGroup α] -- Fuchs p67 -- Bourbaki A.VI.10 Prop 7 @[to_additive] -lemma inf_mul_sup [CovariantClass α α (· * ·) (· ≤ ·)] (a b : α) : (a ⊓ b) * (a ⊔ b) = a * b := +lemma inf_mul_sup [MulLeftMono α] (a b : α) : (a ⊓ b) * (a ⊔ b) = a * b := calc (a ⊓ b) * (a ⊔ b) = (a ⊓ b) * (a * b * (b⁻¹ ⊔ a⁻¹)) := by rw [mul_sup b⁻¹ a⁻¹ (a * b), mul_inv_cancel_right, mul_inv_cancel_comm] @@ -117,7 +117,7 @@ lemma inf_mul_sup [CovariantClass α α (· * ·) (· ≤ ·)] (a b : α) : (a -- Non-comm case needs cancellation law https://ncatlab.org/nlab/show/distributive+lattice @[to_additive "Every lattice ordered commutative additive group is a distributive lattice"] def CommGroup.toDistribLattice (α : Type*) [Lattice α] [CommGroup α] - [CovariantClass α α (· * ·) (· ≤ ·)] : DistribLattice α where + [MulLeftMono α] : DistribLattice α where le_sup_inf x y z := by rw [← mul_le_mul_iff_left (x ⊓ (y ⊓ z)), inf_mul_sup x (y ⊓ z), ← inv_mul_le_iff_le_mul, le_inf_iff] diff --git a/Mathlib/Algebra/Order/Group/MinMax.lean b/Mathlib/Algebra/Order/Group/MinMax.lean index adf6af1c4b904..06f9007ec5e6f 100644 --- a/Mathlib/Algebra/Order/Group/MinMax.lean +++ b/Mathlib/Algebra/Order/Group/MinMax.lean @@ -13,7 +13,7 @@ import Mathlib.Algebra.Order.Monoid.Unbundled.MinMax section -variable {α : Type*} [Group α] [LinearOrder α] [CovariantClass α α (· * ·) (· ≤ ·)] +variable {α : Type*} [Group α] [LinearOrder α] [MulLeftMono α] -- TODO: This duplicates `oneLePart_div_leOnePart` @[to_additive (attr := simp)] @@ -35,14 +35,12 @@ variable {α : Type*} [LinearOrderedCommGroup α] @[to_additive min_neg_neg] theorem min_inv_inv' (a b : α) : min a⁻¹ b⁻¹ = (max a b)⁻¹ := Eq.symm <| (@Monotone.map_max α αᵒᵈ _ _ Inv.inv a b) fun _ _ => - -- Porting note: Explicit `α` necessary to infer `CovariantClass` instance - (@inv_le_inv_iff α _ _ _).mpr + inv_le_inv_iff.mpr @[to_additive max_neg_neg] theorem max_inv_inv' (a b : α) : max a⁻¹ b⁻¹ = (min a b)⁻¹ := Eq.symm <| (@Monotone.map_min α αᵒᵈ _ _ Inv.inv a b) fun _ _ => - -- Porting note: Explicit `α` necessary to infer `CovariantClass` instance - (@inv_le_inv_iff α _ _ _).mpr + inv_le_inv_iff.mpr @[to_additive min_sub_sub_right] theorem min_div_div_right' (a b c : α) : min (a / c) (b / c) = min a b / c := by diff --git a/Mathlib/Algebra/Order/Group/OrderIso.lean b/Mathlib/Algebra/Order/Group/OrderIso.lean index 925163de311e6..21e6893bb45b5 100644 --- a/Mathlib/Algebra/Order/Group/OrderIso.lean +++ b/Mathlib/Algebra/Order/Group/OrderIso.lean @@ -24,8 +24,7 @@ variable [Group α] section TypeclassesLeftRightLE -variable [LE α] [CovariantClass α α (· * ·) (· ≤ ·)] [CovariantClass α α (swap (· * ·)) (· ≤ ·)] - {a b : α} +variable [LE α] [MulLeftMono α] [MulRightMono α] {a b : α} section @@ -71,7 +70,7 @@ variable [Group α] [LE α] section Right -variable [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a : α} +variable [MulRightMono α] {a : α} /-- `Equiv.mulRight` as an `OrderIso`. See also `OrderEmbedding.mulRight`. -/ @[to_additive (attr := simps! (config := { simpRhs := true }) toEquiv apply) @@ -95,7 +94,7 @@ end Right section Left -variable [CovariantClass α α (· * ·) (· ≤ ·)] +variable [MulLeftMono α] /-- `Equiv.mulLeft` as an `OrderIso`. See also `OrderEmbedding.mulLeft`. -/ @[to_additive (attr := simps! (config := { simpRhs := true }) toEquiv apply) diff --git a/Mathlib/Algebra/Order/Group/Pointwise/Bounds.lean b/Mathlib/Algebra/Order/Group/Pointwise/Bounds.lean index 1b1e0d3b70882..d7e523a2a6595 100644 --- a/Mathlib/Algebra/Order/Group/Pointwise/Bounds.lean +++ b/Mathlib/Algebra/Order/Group/Pointwise/Bounds.lean @@ -21,8 +21,8 @@ open scoped Pointwise variable {ι G M : Type*} section Mul -variable [Mul M] [Preorder M] [CovariantClass M M (· * ·) (· ≤ ·)] - [CovariantClass M M (swap (· * ·)) (· ≤ ·)] {f g : ι → M} {s t : Set M} {a b : M} +variable [Mul M] [Preorder M] [MulLeftMono M] + [MulRightMono M] {f g : ι → M} {s t : Set M} {a b : M} @[to_additive] lemma mul_mem_upperBounds_mul (ha : a ∈ upperBounds s) (hb : b ∈ upperBounds t) : @@ -61,8 +61,8 @@ lemma BddBelow.range_mul (hf : BddBelow (range f)) (hg : BddBelow (range g)) : end Mul section InvNeg -variable [Group G] [Preorder G] [CovariantClass G G (· * ·) (· ≤ ·)] - [CovariantClass G G (swap (· * ·)) (· ≤ ·)] {s : Set G} {a : G} +variable [Group G] [Preorder G] [MulLeftMono G] + [MulRightMono G] {s : Set G} {a : G} @[to_additive (attr := simp)] theorem bddAbove_inv : BddAbove s⁻¹ ↔ BddBelow s := diff --git a/Mathlib/Algebra/Order/Group/Pointwise/CompleteLattice.lean b/Mathlib/Algebra/Order/Group/Pointwise/CompleteLattice.lean index f2539b43a96df..64ef925d173b5 100644 --- a/Mathlib/Algebra/Order/Group/Pointwise/CompleteLattice.lean +++ b/Mathlib/Algebra/Order/Group/Pointwise/CompleteLattice.lean @@ -34,7 +34,7 @@ variable [One M] end One section Group -variable [Group M] [CovariantClass M M (· * ·) (· ≤ ·)] [CovariantClass M M (swap (· * ·)) (· ≤ ·)] +variable [Group M] [MulLeftMono M] [MulRightMono M] {s t : Set M} @[to_additive] @@ -84,7 +84,7 @@ variable [One M] end One section Group -variable [Group M] [CovariantClass M M (· * ·) (· ≤ ·)] [CovariantClass M M (swap (· * ·)) (· ≤ ·)] +variable [Group M] [MulLeftMono M] [MulRightMono M] (s t : Set M) @[to_additive] diff --git a/Mathlib/Algebra/Order/Group/PosPart.lean b/Mathlib/Algebra/Order/Group/PosPart.lean index eb8835fa80134..3fae44d631e4e 100644 --- a/Mathlib/Algebra/Order/Group/PosPart.lean +++ b/Mathlib/Algebra/Order/Group/PosPart.lean @@ -108,8 +108,8 @@ lemma leOnePart_le_one' : a⁻ᵐ ≤ 1 ↔ a⁻¹ ≤ 1 := by simp [leOnePart] @[to_additive (attr := simp)] lemma leOnePart_inv (a : α) : a⁻¹⁻ᵐ = a⁺ᵐ := by simp [oneLePart, leOnePart] -section covariantmul -variable [CovariantClass α α (· * ·) (· ≤ ·)] +section MulLeftMono +variable [MulLeftMono α] @[to_additive (attr := simp)] lemma leOnePart_eq_inv : a⁻ᵐ = a⁻¹ ↔ a ≤ 1 := by simp [leOnePart] @@ -140,8 +140,8 @@ lemma oneLePart_leOnePart_injective : Injective fun a : α ↦ (a⁺ᵐ, a⁻ᵐ lemma oneLePart_leOnePart_inj : a⁺ᵐ = b⁺ᵐ ∧ a⁻ᵐ = b⁻ᵐ ↔ a = b := Prod.mk.inj_iff.symm.trans oneLePart_leOnePart_injective.eq_iff -section covariantmulop -variable [CovariantClass α α (swap (· * ·)) (· ≤ ·)] +section MulRightMono +variable [MulRightMono α] @[to_additive] lemma leOnePart_anti : Antitone (leOnePart : α → α) := fun _a _b hab ↦ sup_le_sup_right (inv_le_inv_iff.2 hab) _ @@ -167,14 +167,14 @@ lemma leOnePart_eq_inv_inf_one (a : α) : a⁻ᵐ = (a ⊓ 1)⁻¹ := by rw [← mul_left_inj a⁻ᵐ⁻¹, inf_mul, one_mul, mul_inv_cancel, ← div_eq_mul_inv, oneLePart_div_leOnePart, leOnePart_eq_inv_inf_one, inv_inv] -end covariantmulop +end MulRightMono -end covariantmul +end MulLeftMono end Group section CommGroup -variable [CommGroup α] [CovariantClass α α (· * ·) (· ≤ ·)] +variable [CommGroup α] [MulLeftMono α] -- Bourbaki A.VI.12 (with a and b swapped) @[to_additive] lemma sup_eq_mul_oneLePart_div (a b : α) : a ⊔ b = b * (a / b)⁺ᵐ := by @@ -225,7 +225,7 @@ lemma oneLePart_of_one_lt_oneLePart (ha : 1 < a⁺ᵐ) : a⁺ᵐ = a := by @[to_additive (attr := simp)] lemma oneLePart_lt : a⁺ᵐ < b ↔ a < b ∧ 1 < b := sup_lt_iff section covariantmul -variable [CovariantClass α α (· * ·) (· ≤ ·)] +variable [MulLeftMono α] @[to_additive] lemma leOnePart_eq_ite : a⁻ᵐ = if a ≤ 1 then a⁻¹ else 1 := by simp_rw [← one_le_inv']; rw [leOnePart_def, ← maxDefault, ← sup_eq_maxDefault]; simp_rw [sup_comm] @@ -233,7 +233,7 @@ variable [CovariantClass α α (· * ·) (· ≤ ·)] @[to_additive (attr := simp) negPart_pos_iff] lemma one_lt_ltOnePart_iff : 1 < a⁻ᵐ ↔ a < 1 := lt_iff_lt_of_le_iff_le <| (one_le_leOnePart _).le_iff_eq.trans leOnePart_eq_one -variable [CovariantClass α α (swap (· * ·)) (· ≤ ·)] +variable [MulRightMono α] @[to_additive (attr := simp)] lemma leOnePart_lt : a⁻ᵐ < b ↔ b⁻¹ < a ∧ 1 < b := sup_lt_iff.trans <| by rw [inv_lt'] diff --git a/Mathlib/Algebra/Order/Group/Unbundled/Abs.lean b/Mathlib/Algebra/Order/Group/Unbundled/Abs.lean index bdbb144c05508..d5bc417fef884 100644 --- a/Mathlib/Algebra/Order/Group/Unbundled/Abs.lean +++ b/Mathlib/Algebra/Order/Group/Unbundled/Abs.lean @@ -70,7 +70,7 @@ def abs.unexpander : Lean.PrettyPrinter.Unexpander @[to_additive] lemma mabs_div_comm (a b : α) : |a / b|ₘ = |b / a|ₘ := by rw [← mabs_inv, inv_div] -variable [CovariantClass α α (· * ·) (· ≤ ·)] +variable [MulLeftMono α] @[to_additive] lemma mabs_of_one_le (h : 1 ≤ a) : |a|ₘ = a := sup_eq_left.2 <| (inv_le_one'.2 h).trans h @@ -89,7 +89,7 @@ attribute [gcongr] abs_le_abs_of_nonneg @[to_additive (attr := simp)] lemma mabs_one : |(1 : α)|ₘ = 1 := mabs_of_one_le le_rfl -variable [CovariantClass α α (swap (· * ·)) (· ≤ ·)] +variable [MulRightMono α] @[to_additive (attr := simp) abs_nonneg] lemma one_le_mabs (a : α) : 1 ≤ |a|ₘ := by apply pow_two_semiclosed _ @@ -102,7 +102,7 @@ variable [CovariantClass α α (swap (· * ·)) (· ≤ ·)] end Group section CommGroup -variable [CommGroup α] [CovariantClass α α (· * ·) (· ≤ ·)] +variable [CommGroup α] [MulLeftMono α] -- Banasiak Proposition 2.12, Zaanen 2nd lecture /-- The absolute value satisfies the triangle inequality. -/ @@ -204,7 +204,7 @@ variable [Group α] [LinearOrder α] {a b : α} @[to_additive] lemma lt_of_mabs_lt : |a|ₘ < b → a < b := (le_mabs_self _).trans_lt -variable [CovariantClass α α (· * ·) (· ≤ ·)] {a b : α} +variable [MulLeftMono α] {a b : α} @[to_additive (attr := simp) abs_pos] lemma one_lt_mabs : 1 < |a|ₘ ↔ a ≠ 1 := by obtain ha | rfl | ha := lt_trichotomy a 1 @@ -228,7 +228,7 @@ variable [CovariantClass α α (· * ·) (· ≤ ·)] {a b : α} @[to_additive] lemma inv_mabs_le_inv (a : α) : |a|ₘ⁻¹ ≤ a⁻¹ := by simpa using inv_mabs_le a⁻¹ -variable [CovariantClass α α (swap (· * ·)) (· ≤ ·)] +variable [MulRightMono α] @[to_additive] lemma mabs_ne_one : |a|ₘ ≠ 1 ↔ a ≠ 1 := (one_le_mabs a).gt_iff_ne.symm.trans one_lt_mabs diff --git a/Mathlib/Algebra/Order/Group/Unbundled/Basic.lean b/Mathlib/Algebra/Order/Group/Unbundled/Basic.lean index 4d59951456a6e..95bd5c74cb92e 100644 --- a/Mathlib/Algebra/Order/Group/Unbundled/Basic.lean +++ b/Mathlib/Algebra/Order/Group/Unbundled/Basic.lean @@ -31,9 +31,9 @@ section Group variable [Group α] -section TypeclassesLeftLE +section MulLeftMono -variable [LE α] [CovariantClass α α (· * ·) (· ≤ ·)] {a b c : α} +variable [LE α] [MulLeftMono α] {a b c : α} /-- Uses `left` co(ntra)variant. -/ @[to_additive (attr := simp) "Uses `left` co(ntra)variant."] @@ -73,11 +73,11 @@ theorem inv_mul_le_one_iff : a⁻¹ * b ≤ 1 ↔ b ≤ a := -- Porting note: why is the `_root_` needed? _root_.trans inv_mul_le_iff_le_mul <| by rw [mul_one] -end TypeclassesLeftLE +end MulLeftMono -section TypeclassesLeftLT +section MulLeftStrictMono -variable [LT α] [CovariantClass α α (· * ·) (· < ·)] {a b c : α} +variable [LT α] [MulLeftStrictMono α] {a b c : α} /-- Uses `left` co(ntra)variant. -/ @[to_additive (attr := simp) Left.neg_pos_iff "Uses `left` co(ntra)variant."] @@ -114,11 +114,11 @@ theorem lt_inv_mul_iff_lt : 1 < b⁻¹ * a ↔ b < a := by theorem inv_mul_lt_one_iff : a⁻¹ * b < 1 ↔ b < a := _root_.trans inv_mul_lt_iff_lt_mul <| by rw [mul_one] -end TypeclassesLeftLT +end MulLeftStrictMono -section TypeclassesRightLE +section MulRightMono -variable [LE α] [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a b c : α} +variable [LE α] [MulRightMono α] {a b c : α} /-- Uses `right` co(ntra)variant. -/ @[to_additive (attr := simp) "Uses `right` co(ntra)variant."] @@ -161,11 +161,11 @@ theorem le_mul_inv_iff_le : 1 ≤ a * b⁻¹ ↔ b ≤ a := by theorem mul_inv_le_one_iff : b * a⁻¹ ≤ 1 ↔ b ≤ a := _root_.trans mul_inv_le_iff_le_mul <| by rw [one_mul] -end TypeclassesRightLE +end MulRightMono -section TypeclassesRightLT +section MulRightStrictMono -variable [LT α] [CovariantClass α α (swap (· * ·)) (· < ·)] {a b c : α} +variable [LT α] [MulRightStrictMono α] {a b c : α} /-- Uses `right` co(ntra)variant. -/ @[to_additive (attr := simp) "Uses `right` co(ntra)variant."] @@ -206,12 +206,11 @@ theorem lt_mul_inv_iff_lt : 1 < a * b⁻¹ ↔ b < a := by theorem mul_inv_lt_one_iff : b * a⁻¹ < 1 ↔ b < a := _root_.trans mul_inv_lt_iff_lt_mul <| by rw [one_mul] -end TypeclassesRightLT +end MulRightStrictMono -section TypeclassesLeftRightLE +section MulLeftMono_MulRightMono -variable [LE α] [CovariantClass α α (· * ·) (· ≤ ·)] - {a b c d : α} +variable [LE α] [MulLeftMono α] {a b c d : α} @[to_additive (attr := simp)] theorem div_le_self_iff (a : α) {b : α} : a / b ≤ a ↔ 1 ≤ b := by @@ -223,7 +222,7 @@ theorem le_div_self_iff (a : α) {b : α} : a ≤ a / b ↔ b ≤ 1 := by alias ⟨_, sub_le_self⟩ := sub_le_self_iff -variable [CovariantClass α α (swap (· * ·)) (· ≤ ·)] +variable [MulRightMono α] @[to_additive (attr := simp)] theorem inv_le_inv_iff : a⁻¹ ≤ b⁻¹ ↔ b ≤ a := by @@ -237,12 +236,11 @@ theorem mul_inv_le_inv_mul_iff : a * b⁻¹ ≤ d⁻¹ * c ↔ d * a ≤ c * b : rw [← mul_le_mul_iff_left d, ← mul_le_mul_iff_right b, mul_inv_cancel_left, mul_assoc, inv_mul_cancel_right] -end TypeclassesLeftRightLE +end MulLeftMono_MulRightMono -section TypeclassesLeftRightLT +section MulLeftStrictMono_MulRightStrictMono -variable [LT α] [CovariantClass α α (· * ·) (· < ·)] - {a b c d : α} +variable [LT α] [MulLeftStrictMono α] {a b c d : α} @[to_additive (attr := simp)] theorem div_lt_self_iff (a : α) {b : α} : a / b < a ↔ 1 < b := by @@ -250,7 +248,7 @@ theorem div_lt_self_iff (a : α) {b : α} : a / b < a ↔ 1 < b := by alias ⟨_, sub_lt_self⟩ := sub_lt_self_iff -variable [CovariantClass α α (swap (· * ·)) (· < ·)] +variable [MulRightStrictMono α] @[to_additive (attr := simp)] theorem inv_lt_inv_iff : a⁻¹ < b⁻¹ ↔ b < a := by @@ -276,7 +274,7 @@ theorem mul_inv_lt_inv_mul_iff : a * b⁻¹ < d⁻¹ * c ↔ d * a < c * b := by rw [← mul_lt_mul_iff_left d, ← mul_lt_mul_iff_right b, mul_inv_cancel_left, mul_assoc, inv_mul_cancel_right] -end TypeclassesLeftRightLT +end MulLeftStrictMono_MulRightStrictMono section Preorder @@ -284,7 +282,7 @@ variable [Preorder α] section LeftLE -variable [CovariantClass α α (· * ·) (· ≤ ·)] {a : α} +variable [MulLeftMono α] {a : α} @[to_additive] theorem Left.inv_le_self (h : 1 ≤ a) : a⁻¹ ≤ a := @@ -300,7 +298,7 @@ end LeftLE section LeftLT -variable [CovariantClass α α (· * ·) (· < ·)] {a : α} +variable [MulLeftStrictMono α] {a : α} @[to_additive] theorem Left.inv_lt_self (h : 1 < a) : a⁻¹ < a := @@ -316,7 +314,7 @@ end LeftLT section RightLE -variable [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a : α} +variable [MulRightMono α] {a : α} @[to_additive] theorem Right.inv_le_self (h : 1 ≤ a) : a⁻¹ ≤ a := @@ -330,7 +328,7 @@ end RightLE section RightLT -variable [CovariantClass α α (swap (· * ·)) (· < ·)] {a : α} +variable [MulRightStrictMono α] {a : α} @[to_additive] theorem Right.inv_lt_self (h : 1 < a) : a⁻¹ < a := @@ -352,7 +350,7 @@ variable [CommGroup α] section LE -variable [LE α] [CovariantClass α α (· * ·) (· ≤ ·)] {a b c d : α} +variable [LE α] [MulLeftMono α] {a b c d : α} @[to_additive] theorem inv_mul_le_iff_le_mul' : c⁻¹ * a ≤ b ↔ a ≤ b * c := by rw [inv_mul_le_iff_le_mul, mul_comm] @@ -370,7 +368,7 @@ end LE section LT -variable [LT α] [CovariantClass α α (· * ·) (· < ·)] {a b c d : α} +variable [LT α] [MulLeftStrictMono α] {a b c d : α} @[to_additive] theorem inv_mul_lt_iff_lt_mul' : c⁻¹ * a < b ↔ a < b * c := by rw [inv_mul_lt_iff_lt_mul, mul_comm] @@ -471,7 +469,7 @@ variable [Group α] [LE α] section Right -variable [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a b c : α} +variable [MulRightMono α] {a b c : α} @[to_additive] theorem div_le_div_iff_right (c : α) : a / c ≤ b / c ↔ a ≤ b := by @@ -511,15 +509,14 @@ attribute [simp] div_le_iff_le_mul -- (a renamed version of) `tsub_le_iff_right`? -- see Note [lower instance priority] instance (priority := 100) AddGroup.toOrderedSub {α : Type*} [AddGroup α] [LE α] - [CovariantClass α α (swap (· + ·)) (· ≤ ·)] : OrderedSub α := + [AddRightMono α] : OrderedSub α := ⟨fun _ _ _ => sub_le_iff_le_add⟩ end Right section Left -variable [CovariantClass α α (· * ·) (· ≤ ·)] -variable [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a b c : α} +variable [MulLeftMono α] [MulRightMono α] {a b c : α} @[to_additive] theorem div_le_div_iff_left (a : α) : a / b ≤ a / c ↔ c ≤ b := by @@ -540,7 +537,7 @@ variable [CommGroup α] section LE -variable [LE α] [CovariantClass α α (· * ·) (· ≤ ·)] {a b c d : α} +variable [LE α] [MulLeftMono α] {a b c d : α} @[to_additive sub_le_sub_iff] theorem div_le_div_iff' : a / b ≤ c / d ↔ a * d ≤ c * b := by @@ -575,7 +572,7 @@ end LE section Preorder -variable [Preorder α] [CovariantClass α α (· * ·) (· ≤ ·)] {a b c d : α} +variable [Preorder α] [MulLeftMono α] {a b c d : α} @[to_additive (attr := gcongr) sub_le_sub] theorem div_le_div'' (hab : a ≤ b) (hcd : c ≤ d) : a / d ≤ b / c := by @@ -594,7 +591,7 @@ variable [Group α] [LT α] section Right -variable [CovariantClass α α (swap (· * ·)) (· < ·)] {a b c : α} +variable [MulRightStrictMono α] {a b c : α} @[to_additive (attr := simp)] theorem div_lt_div_iff_right (c : α) : a / c < b / c ↔ a < b := by @@ -634,7 +631,7 @@ end Right section Left -variable [CovariantClass α α (· * ·) (· < ·)] [CovariantClass α α (swap (· * ·)) (· < ·)] +variable [MulLeftStrictMono α] [MulRightStrictMono α] {a b c : α} @[to_additive (attr := simp)] @@ -660,7 +657,7 @@ variable [CommGroup α] section LT -variable [LT α] [CovariantClass α α (· * ·) (· < ·)] {a b c d : α} +variable [LT α] [MulLeftStrictMono α] {a b c d : α} @[to_additive sub_lt_sub_iff] theorem div_lt_div_iff' : a / b < c / d ↔ a * d < c * b := by @@ -692,7 +689,7 @@ end LT section Preorder -variable [Preorder α] [CovariantClass α α (· * ·) (· < ·)] {a b c d : α} +variable [Preorder α] [MulLeftStrictMono α] {a b c d : α} @[to_additive (attr := gcongr) sub_lt_sub] theorem div_lt_div'' (hab : a < b) (hcd : c < d) : a / d < b / c := by @@ -702,7 +699,7 @@ theorem div_lt_div'' (hab : a < b) (hcd : c < d) : a / d < b / c := by end Preorder section LinearOrder -variable [LinearOrder α] [CovariantClass α α (· * ·) (· ≤ ·)] {a b c d : α} +variable [LinearOrder α] [MulLeftMono α] {a b c d : α} @[to_additive] lemma lt_or_lt_of_div_lt_div : a / d < b / c → a < b ∨ c < d := by contrapose!; exact fun h ↦ div_le_div'' h.1 h.2 @@ -715,10 +712,10 @@ section LinearOrder variable [Group α] [LinearOrder α] @[to_additive (attr := simp) cmp_sub_zero] -theorem cmp_div_one' [CovariantClass α α (swap (· * ·)) (· ≤ ·)] (a b : α) : +theorem cmp_div_one' [MulRightMono α] (a b : α) : cmp (a / b) 1 = cmp a b := by rw [← cmp_mul_right' _ _ b, one_mul, div_mul_cancel] -variable [CovariantClass α α (· * ·) (· ≤ ·)] +variable [MulLeftMono α] section VariableNames @@ -735,7 +732,7 @@ theorem le_iff_forall_one_lt_lt_mul : a ≤ b ↔ ∀ ε, 1 < ε → a < b * ε /- I (DT) introduced this lemma to prove (the additive version `sub_le_sub_flip` of) `div_le_div_flip` below. Now I wonder what is the point of either of these lemmas... -/ @[to_additive] -theorem div_le_inv_mul_iff [CovariantClass α α (swap (· * ·)) (· ≤ ·)] : +theorem div_le_inv_mul_iff [MulRightMono α] : a / b ≤ a⁻¹ * b ↔ a ≤ b := by rw [div_eq_mul_inv, mul_inv_le_inv_mul_iff] exact @@ -747,7 +744,7 @@ theorem div_le_inv_mul_iff [CovariantClass α α (swap (· * ·)) (· ≤ ·)] : -- since the LHS simplifies with `tsub_le_iff_right` @[to_additive] theorem div_le_div_flip {α : Type*} [CommGroup α] [LinearOrder α] - [CovariantClass α α (· * ·) (· ≤ ·)] {a b : α} : a / b ≤ b / a ↔ a ≤ b := by + [MulLeftMono α] {a b : α} : a / b ≤ b / a ↔ a ≤ b := by rw [div_eq_mul_inv b, mul_comm] exact div_le_inv_mul_iff @@ -757,8 +754,8 @@ end LinearOrder section -variable {β : Type*} [Group α] [Preorder α] [CovariantClass α α (· * ·) (· ≤ ·)] - [CovariantClass α α (swap (· * ·)) (· ≤ ·)] [Preorder β] {f : β → α} {s : Set β} +variable {β : Type*} [Group α] [Preorder α] [MulLeftMono α] + [MulRightMono α] [Preorder β] {f : β → α} {s : Set β} @[to_additive] theorem Monotone.inv (hf : Monotone f) : Antitone fun x => (f x)⁻¹ := fun _ _ hxy => @@ -780,8 +777,8 @@ end section -variable {β : Type*} [Group α] [Preorder α] [CovariantClass α α (· * ·) (· < ·)] - [CovariantClass α α (swap (· * ·)) (· < ·)] [Preorder β] {f : β → α} {s : Set β} +variable {β : Type*} [Group α] [Preorder α] [MulLeftStrictMono α] + [MulRightStrictMono α] [Preorder β] {f : β → α} {s : Set β} @[to_additive] theorem StrictMono.inv (hf : StrictMono f) : StrictAnti fun x => (f x)⁻¹ := fun _ _ hxy => diff --git a/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean b/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean index 930a761c341ac..b0a55e98d4250 100644 --- a/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean +++ b/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean @@ -291,8 +291,8 @@ lemma zero_eq_bot : (0 : WithZero α) = ⊥ := rfl theorem coe_le_iff {x : WithZero α} : (a : WithZero α) ≤ x ↔ ∃ b : α, x = b ∧ a ≤ b := WithBot.coe_le_iff -instance covariantClass_mul_le [Mul α] [CovariantClass α α (· * ·) (· ≤ ·)] : - CovariantClass (WithZero α) (WithZero α) (· * ·) (· ≤ ·) := by +instance mulLeftMono [Mul α] [MulLeftMono α] : + MulLeftMono (WithZero α) := by refine ⟨fun a b c hbc => ?_⟩ induction a; · exact zero_le _ induction b; · exact zero_le _ @@ -300,8 +300,8 @@ instance covariantClass_mul_le [Mul α] [CovariantClass α α (· * ·) (· ≤ rw [← coe_mul _ c, ← coe_mul, coe_le_coe] exact mul_le_mul_left' hbc' _ -protected lemma covariantClass_add_le [AddZeroClass α] [CovariantClass α α (· + ·) (· ≤ ·)] - (h : ∀ a : α, 0 ≤ a) : CovariantClass (WithZero α) (WithZero α) (· + ·) (· ≤ ·) := by +protected lemma addLeftMono [AddZeroClass α] [AddLeftMono α] + (h : ∀ a : α, 0 ≤ a) : AddLeftMono (WithZero α) := by refine ⟨fun a b c hbc => ?_⟩ induction a · rwa [zero_add, zero_add] @@ -332,8 +332,8 @@ variable [PartialOrder α] instance partialOrder : PartialOrder (WithZero α) := WithBot.partialOrder -instance contravariantClass_mul_lt [Mul α] [ContravariantClass α α (· * ·) (· < ·)] : - ContravariantClass (WithZero α) (WithZero α) (· * ·) (· < ·) := by +instance mulLeftReflectLT [Mul α] [MulLeftReflectLT α] : + MulLeftReflectLT (WithZero α) := by refine ⟨fun a b c h => ?_⟩ have := ((zero_le _).trans_lt h).ne' induction a @@ -378,7 +378,7 @@ elements are ≤ 1 and then 1 is the top element. protected abbrev orderedAddCommMonoid [OrderedAddCommMonoid α] (zero_le : ∀ a : α, 0 ≤ a) : OrderedAddCommMonoid (WithZero α) := { WithZero.partialOrder, WithZero.addCommMonoid with - add_le_add_left := @add_le_add_left _ _ _ (WithZero.covariantClass_add_le zero_le).. } + add_le_add_left := @add_le_add_left _ _ _ (WithZero.addLeftMono zero_le).. } -- This instance looks absurd: a monoid already has a zero /-- Adding a new zero to a canonically ordered additive monoid produces another one. -/ diff --git a/Mathlib/Algebra/Order/GroupWithZero/WithZero.lean b/Mathlib/Algebra/Order/GroupWithZero/WithZero.lean index d5b2da40325ae..1595fba036e41 100644 --- a/Mathlib/Algebra/Order/GroupWithZero/WithZero.lean +++ b/Mathlib/Algebra/Order/GroupWithZero/WithZero.lean @@ -28,7 +28,7 @@ theory. These instances enable lemmas such as `mul_pos` to fire on `ℤₘ₀`. assert_not_exists Ring -- this makes `mul_lt_mul_left`, `mul_pos` etc work on `ℤₘ₀` -instance {α : Type*} [Mul α] [Preorder α] [CovariantClass α α (· * ·) (· < ·)] : +instance {α : Type*} [Mul α] [Preorder α] [MulLeftStrictMono α] : PosMulStrictMono (WithZero α) where elim := @fun | ⟨(x : α), hx⟩, 0, (b : α), _ => by @@ -39,7 +39,7 @@ instance {α : Type*} [Mul α] [Preorder α] [CovariantClass α α (· * ·) (· exact mul_lt_mul_left' h x open Function in -instance {α : Type*} [Mul α] [Preorder α] [CovariantClass α α (swap (· * ·)) (· < ·)] : +instance {α : Type*} [Mul α] [Preorder α] [MulRightStrictMono α] : MulPosStrictMono (WithZero α) where elim := @fun | ⟨(x : α), hx⟩, 0, (b : α), _ => by @@ -49,7 +49,7 @@ instance {α : Type*} [Mul α] [Preorder α] [CovariantClass α α (swap (· * norm_cast at h ⊢ exact mul_lt_mul_right' h x -instance {α : Type*} [Mul α] [Preorder α] [CovariantClass α α (· * ·) (· ≤ ·)] : +instance {α : Type*} [Mul α] [Preorder α] [MulLeftMono α] : PosMulMono (WithZero α) where elim := @fun | ⟨0, _⟩, a, b, _ => by @@ -65,7 +65,7 @@ instance {α : Type*} [Mul α] [Preorder α] [CovariantClass α α (· * ·) (· -- This makes `lt_mul_of_le_of_one_lt'` work on `ℤₘ₀` open Function in -instance {α : Type*} [Mul α] [Preorder α] [CovariantClass α α (swap (· * ·)) (· ≤ ·)] : +instance {α : Type*} [Mul α] [Preorder α] [MulRightMono α] : MulPosMono (WithZero α) where elim := @fun | ⟨0, _⟩, a, b, _ => by diff --git a/Mathlib/Algebra/Order/Hom/Monoid.lean b/Mathlib/Algebra/Order/Hom/Monoid.lean index 89356b58c0ecf..2c821fe45eb5c 100644 --- a/Mathlib/Algebra/Order/Hom/Monoid.lean +++ b/Mathlib/Algebra/Order/Hom/Monoid.lean @@ -265,7 +265,7 @@ theorem monotone_iff_map_nonpos : Monotone (f : α → β) ↔ ∀ a ≤ 0, f a theorem antitone_iff_map_nonneg : Antitone (f : α → β) ↔ ∀ a ≤ 0, 0 ≤ f a := monotone_comp_ofDual_iff.symm.trans <| monotone_iff_map_nonneg (α := αᵒᵈ) (iamhc := iamhc) _ -variable [CovariantClass β β (· + ·) (· < ·)] +variable [AddLeftStrictMono β] theorem strictMono_iff_map_pos : StrictMono (f : α → β) ↔ ∀ a, 0 < a → 0 < f a := by diff --git a/Mathlib/Algebra/Order/Interval/Basic.lean b/Mathlib/Algebra/Order/Interval/Basic.lean index bcfec00e2cd85..a8d682494ed1f 100644 --- a/Mathlib/Algebra/Order/Interval/Basic.lean +++ b/Mathlib/Algebra/Order/Interval/Basic.lean @@ -118,8 +118,7 @@ Note that this multiplication does not apply to `ℚ` or `ℝ`. section Mul -variable [Preorder α] [Mul α] [CovariantClass α α (· * ·) (· ≤ ·)] - [CovariantClass α α (swap (· * ·)) (· ≤ ·)] +variable [Preorder α] [Mul α] [MulLeftMono α] [MulRightMono α] @[to_additive] instance : Mul (NonemptyInterval α) := @@ -178,8 +177,8 @@ end Mul -- TODO: if `to_additive` gets improved sufficiently, derive this from `hasPow` -instance NonemptyInterval.hasNSMul [AddMonoid α] [Preorder α] [CovariantClass α α (· + ·) (· ≤ ·)] - [CovariantClass α α (swap (· + ·)) (· ≤ ·)] : SMul ℕ (NonemptyInterval α) := +instance NonemptyInterval.hasNSMul [AddMonoid α] [Preorder α] [AddLeftMono α] + [AddRightMono α] : SMul ℕ (NonemptyInterval α) := ⟨fun n s => ⟨(n • s.fst, n • s.snd), nsmul_le_nsmul_right s.fst_le_snd _⟩⟩ section Pow @@ -187,14 +186,13 @@ section Pow variable [Monoid α] [Preorder α] @[to_additive existing] -instance NonemptyInterval.hasPow - [CovariantClass α α (· * ·) (· ≤ ·)] [CovariantClass α α (swap (· * ·)) (· ≤ ·)] : +instance NonemptyInterval.hasPow [MulLeftMono α] [MulRightMono α] : Pow (NonemptyInterval α) ℕ := ⟨fun s n => ⟨s.toProd ^ n, pow_le_pow_left' s.fst_le_snd _⟩⟩ namespace NonemptyInterval -variable [CovariantClass α α (· * ·) (· ≤ ·)] [CovariantClass α α (swap (· * ·)) (· ≤ ·)] +variable [MulLeftMono α] [MulRightMono α] variable (s : NonemptyInterval α) (a : α) (n : ℕ) @[to_additive (attr := simp) toProd_nsmul] @@ -275,8 +273,7 @@ is not a thing and probably should not become one). section Sub -variable [Preorder α] [AddCommSemigroup α] [Sub α] [OrderedSub α] - [CovariantClass α α (· + ·) (· ≤ ·)] +variable [Preorder α] [AddCommSemigroup α] [Sub α] [OrderedSub α] [AddLeftMono α] instance : Sub (NonemptyInterval α) := ⟨fun s t => ⟨(s.fst - t.snd, s.snd - t.fst), tsub_le_tsub s.fst_le_snd t.fst_le_snd⟩⟩ @@ -334,7 +331,7 @@ Note that this division does not apply to `ℚ` or `ℝ`. section Div -variable [Preorder α] [CommGroup α] [CovariantClass α α (· * ·) (· ≤ ·)] +variable [Preorder α] [CommGroup α] [MulLeftMono α] @[to_additive existing] instance : Div (NonemptyInterval α) := diff --git a/Mathlib/Algebra/Order/Kleene.lean b/Mathlib/Algebra/Order/Kleene.lean index 08c772b601a9e..bc8e1a8feedb4 100644 --- a/Mathlib/Algebra/Order/Kleene.lean +++ b/Mathlib/Algebra/Order/Kleene.lean @@ -162,13 +162,11 @@ instance (priority := 100) IdemSemiring.toCanonicallyOrderedAddCommMonoid : le_self_add := fun a b ↦ add_eq_right_iff_le.1 <| by rw [← add_assoc, add_idem] } -- See note [lower instance priority] -instance (priority := 100) IdemSemiring.toCovariantClass_mul_le : - CovariantClass α α (· * ·) (· ≤ ·) := +instance (priority := 100) IdemSemiring.toMulLeftMono : MulLeftMono α := ⟨fun a b c hbc ↦ add_eq_left_iff_le.1 <| by rw [← mul_add, hbc.add_eq_left]⟩ -- See note [lower instance priority] -instance (priority := 100) IdemSemiring.toCovariantClass_swap_mul_le : - CovariantClass α α (swap (· * ·)) (· ≤ ·) := +instance (priority := 100) IdemSemiring.toMulRightMono : MulRightMono α := ⟨fun a b c hbc ↦ add_eq_left_iff_le.1 <| by rw [← add_mul, hbc.add_eq_left]⟩ end IdemSemiring diff --git a/Mathlib/Algebra/Order/Module/OrderedSMul.lean b/Mathlib/Algebra/Order/Module/OrderedSMul.lean index ff7892c724901..6afca816cb005 100644 --- a/Mathlib/Algebra/Order/Module/OrderedSMul.lean +++ b/Mathlib/Algebra/Order/Module/OrderedSMul.lean @@ -136,7 +136,7 @@ end LinearOrderedSemifield section Invertible variable (α : Type*) {β : Type*} variable [Semiring α] [Invertible (2 : α)] [Lattice β] [AddCommGroup β] [Module α β] - [CovariantClass β β (· + ·) (· ≤ ·)] + [AddLeftMono β] lemma inf_eq_half_smul_add_sub_abs_sub (x y : β) : x ⊓ y = (⅟2 : α) • (x + y - |y - x|) := by rw [← two_nsmul_inf_eq_add_sub_abs_sub x y, two_smul, ← two_smul α, @@ -151,7 +151,7 @@ end Invertible section DivisionSemiring variable (α : Type*) {β : Type*} variable [DivisionSemiring α] [NeZero (2 : α)] [Lattice β] [AddCommGroup β] [Module α β] - [CovariantClass β β (· + ·) (· ≤ ·)] + [AddLeftMono β] lemma inf_eq_half_smul_add_sub_abs_sub' (x y : β) : x ⊓ y = (2⁻¹ : α) • (x + y - |y - x|) := by letI := invertibleOfNonzero (two_ne_zero' α) diff --git a/Mathlib/Algebra/Order/Monoid/Basic.lean b/Mathlib/Algebra/Order/Monoid/Basic.lean index 824c12d19dddf..32184a47b1ed6 100644 --- a/Mathlib/Algebra/Order/Monoid/Basic.lean +++ b/Mathlib/Algebra/Order/Monoid/Basic.lean @@ -73,7 +73,7 @@ See also `OrderIso.mulLeft` when working in an ordered group. -/ "The order embedding sending `b` to `a + b`, for some fixed `a`. See also `OrderIso.addLeft` when working in an additive ordered group."] def OrderEmbedding.mulLeft {α : Type*} [Mul α] [LinearOrder α] - [CovariantClass α α (· * ·) (· < ·)] (m : α) : α ↪o α := + [MulLeftStrictMono α] (m : α) : α ↪o α := OrderEmbedding.ofStrictMono (fun n => m * n) fun _ _ w => mul_lt_mul_left' w m /-- The order embedding sending `b` to `b * a`, for some fixed `a`. @@ -82,5 +82,5 @@ See also `OrderIso.mulRight` when working in an ordered group. -/ "The order embedding sending `b` to `b + a`, for some fixed `a`. See also `OrderIso.addRight` when working in an additive ordered group."] def OrderEmbedding.mulRight {α : Type*} [Mul α] [LinearOrder α] - [CovariantClass α α (swap (· * ·)) (· < ·)] (m : α) : α ↪o α := + [MulRightStrictMono α] (m : α) : α ↪o α := OrderEmbedding.ofStrictMono (fun n => n * m) fun _ _ w => mul_lt_mul_right' w m diff --git a/Mathlib/Algebra/Order/Monoid/Canonical/Defs.lean b/Mathlib/Algebra/Order/Monoid/Canonical/Defs.lean index 3a2da1f5ad239..f3c8c9e9daa85 100644 --- a/Mathlib/Algebra/Order/Monoid/Canonical/Defs.lean +++ b/Mathlib/Algebra/Order/Monoid/Canonical/Defs.lean @@ -156,7 +156,7 @@ theorem le_mul_right (h : a ≤ b) : a ≤ b * c := _ ≤ b * c := mul_le_mul' h (one_le _) @[to_additive] -theorem lt_iff_exists_mul [CovariantClass α α (· * ·) (· < ·)] : a < b ↔ ∃ c > 1, b = a * c := by +theorem lt_iff_exists_mul [MulLeftStrictMono α] : a < b ↔ ∃ c > 1, b = a * c := by rw [lt_iff_le_and_ne, le_iff_exists_mul, ← exists_and_right] apply exists_congr intro c diff --git a/Mathlib/Algebra/Order/Monoid/Defs.lean b/Mathlib/Algebra/Order/Monoid/Defs.lean index c2b31a4a0ca96..32ed5678901e7 100644 --- a/Mathlib/Algebra/Order/Monoid/Defs.lean +++ b/Mathlib/Algebra/Order/Monoid/Defs.lean @@ -32,12 +32,12 @@ section OrderedCommMonoid variable [OrderedCommMonoid α] @[to_additive] -instance OrderedCommMonoid.toCovariantClassLeft : CovariantClass α α (· * ·) (· ≤ ·) where +instance OrderedCommMonoid.toMulLeftMono : MulLeftMono α where elim := fun a _ _ bc ↦ OrderedCommMonoid.mul_le_mul_left _ _ bc a @[to_additive] -theorem OrderedCommMonoid.toCovariantClassRight (M : Type*) [OrderedCommMonoid M] : - CovariantClass M M (swap (· * ·)) (· ≤ ·) := +theorem OrderedCommMonoid.toMulRightMono (M : Type*) [OrderedCommMonoid M] : + MulRightMono M := inferInstance end OrderedCommMonoid @@ -58,18 +58,18 @@ variable [OrderedCancelCommMonoid α] -- See note [lower instance priority] @[to_additive] -instance (priority := 200) OrderedCancelCommMonoid.toContravariantClassLeLeft : - ContravariantClass α α (· * ·) (· ≤ ·) := +instance (priority := 200) OrderedCancelCommMonoid.toMulLeftReflectLE : + MulLeftReflectLE α := ⟨OrderedCancelCommMonoid.le_of_mul_le_mul_left⟩ @[to_additive] -instance OrderedCancelCommMonoid.toContravariantClassLeft : - ContravariantClass α α (· * ·) (· < ·) where +instance OrderedCancelCommMonoid.toMulLeftReflectLT : + MulLeftReflectLT α where elim := contravariant_lt_of_contravariant_le α α _ ContravariantClass.elim @[to_additive] -theorem OrderedCancelCommMonoid.toContravariantClassRight : - ContravariantClass α α (swap (· * ·)) (· < ·) := +theorem OrderedCancelCommMonoid.toMulRightReflectLT : + MulRightReflectLT α := inferInstance -- See note [lower instance priority] diff --git a/Mathlib/Algebra/Order/Monoid/NatCast.lean b/Mathlib/Algebra/Order/Monoid/NatCast.lean index f5a1d382ff7a2..d1d107db41833 100644 --- a/Mathlib/Algebra/Order/Monoid/NatCast.lean +++ b/Mathlib/Algebra/Order/Monoid/NatCast.lean @@ -16,37 +16,37 @@ variable {α : Type*} open Function lemma lt_add_one [One α] [AddZeroClass α] [PartialOrder α] [ZeroLEOneClass α] - [NeZero (1 : α)] [CovariantClass α α (·+·) (·<·)] (a : α) : a < a + 1 := + [NeZero (1 : α)] [AddLeftStrictMono α] (a : α) : a < a + 1 := lt_add_of_pos_right _ zero_lt_one lemma lt_one_add [One α] [AddZeroClass α] [PartialOrder α] [ZeroLEOneClass α] - [NeZero (1 : α)] [CovariantClass α α (swap (·+·)) (·<·)] (a : α) : a < 1 + a := + [NeZero (1 : α)] [AddRightStrictMono α] (a : α) : a < 1 + a := lt_add_of_pos_left _ zero_lt_one variable [AddMonoidWithOne α] -lemma zero_le_two [Preorder α] [ZeroLEOneClass α] [CovariantClass α α (·+·) (·≤·)] : +lemma zero_le_two [Preorder α] [ZeroLEOneClass α] [AddLeftMono α] : (0 : α) ≤ 2 := by rw [← one_add_one_eq_two] exact add_nonneg zero_le_one zero_le_one -lemma zero_le_three [Preorder α] [ZeroLEOneClass α] [CovariantClass α α (·+·) (·≤·)] : +lemma zero_le_three [Preorder α] [ZeroLEOneClass α] [AddLeftMono α] : (0 : α) ≤ 3 := by rw [← two_add_one_eq_three] exact add_nonneg zero_le_two zero_le_one -lemma zero_le_four [Preorder α] [ZeroLEOneClass α] [CovariantClass α α (·+·) (·≤·)] : +lemma zero_le_four [Preorder α] [ZeroLEOneClass α] [AddLeftMono α] : (0 : α) ≤ 4 := by rw [← three_add_one_eq_four] exact add_nonneg zero_le_three zero_le_one -lemma one_le_two [LE α] [ZeroLEOneClass α] [CovariantClass α α (·+·) (·≤·)] : +lemma one_le_two [LE α] [ZeroLEOneClass α] [AddLeftMono α] : (1 : α) ≤ 2 := calc (1 : α) = 1 + 0 := (add_zero 1).symm _ ≤ 1 + 1 := add_le_add_left zero_le_one _ _ = 2 := one_add_one_eq_two -lemma one_le_two' [LE α] [ZeroLEOneClass α] [CovariantClass α α (swap (·+·)) (·≤·)] : +lemma one_le_two' [LE α] [ZeroLEOneClass α] [AddRightMono α] : (1 : α) ≤ 2 := calc (1 : α) = 0 + 1 := (zero_add 1).symm _ ≤ 1 + 1 := add_le_add_right zero_le_one _ @@ -56,7 +56,7 @@ section variable [PartialOrder α] [ZeroLEOneClass α] [NeZero (1 : α)] section -variable [CovariantClass α α (·+·) (·≤·)] +variable [AddLeftMono α] /-- See `zero_lt_two'` for a version with the type explicit. -/ @[simp] lemma zero_lt_two : (0 : α) < 2 := zero_lt_one.trans_le one_le_two @@ -88,7 +88,7 @@ instance ZeroLEOneClass.neZero.four : NeZero (4 : α) := ⟨zero_lt_four.ne'⟩ end -lemma one_lt_two [CovariantClass α α (·+·) (·<·)] : (1 : α) < 2 := by +lemma one_lt_two [AddLeftStrictMono α] : (1 : α) < 2 := by rw [← one_add_one_eq_two] exact lt_add_one _ diff --git a/Mathlib/Algebra/Order/Monoid/OrderDual.lean b/Mathlib/Algebra/Order/Monoid/OrderDual.lean index 6c9ac68a97316..a56f9ffa0a23b 100644 --- a/Mathlib/Algebra/Order/Monoid/OrderDual.lean +++ b/Mathlib/Algebra/Order/Monoid/OrderDual.lean @@ -21,9 +21,9 @@ namespace OrderDual instance orderedCommMonoid [OrderedCommMonoid α] : OrderedCommMonoid αᵒᵈ := { mul_le_mul_left := fun _ _ h c => mul_le_mul_left' h c } -@[to_additive OrderDual.OrderedCancelAddCommMonoid.to_contravariantClass] -instance OrderedCancelCommMonoid.to_contravariantClass [OrderedCancelCommMonoid α] : - ContravariantClass αᵒᵈ αᵒᵈ HMul.hMul LE.le where +@[to_additive OrderDual.OrderedCancelAddCommMonoid.to_mulLeftReflectLE] +instance OrderedCancelCommMonoid.to_mulLeftReflectLE [OrderedCancelCommMonoid α] : + MulLeftReflectLE αᵒᵈ where elim a b c := OrderedCancelCommMonoid.le_of_mul_le_mul_left (α := α) a c b -- Porting note: Lean 3 to_additive name omits first namespace part diff --git a/Mathlib/Algebra/Order/Monoid/Prod.lean b/Mathlib/Algebra/Order/Monoid/Prod.lean index e3d9ab5369849..4ed2d8ca1fa79 100644 --- a/Mathlib/Algebra/Order/Monoid/Prod.lean +++ b/Mathlib/Algebra/Order/Monoid/Prod.lean @@ -46,7 +46,7 @@ namespace Lex @[to_additive] instance orderedCommMonoid [OrderedCommMonoid α] - [CovariantClass α α (· * ·) (· < ·)] [OrderedCommMonoid β] : + [MulLeftStrictMono α] [OrderedCommMonoid β] : OrderedCommMonoid (α ×ₗ β) where mul_le_mul_left _ _ hxy z := ((le_iff _ _).1 hxy).elim (fun hxy => left _ _ <| mul_lt_mul_left' hxy _) diff --git a/Mathlib/Algebra/Order/Monoid/Submonoid.lean b/Mathlib/Algebra/Order/Monoid/Submonoid.lean index d88dd033520c1..7dd9498120011 100644 --- a/Mathlib/Algebra/Order/Monoid/Submonoid.lean +++ b/Mathlib/Algebra/Order/Monoid/Submonoid.lean @@ -88,7 +88,7 @@ instance toLinearOrderedCancelCommMonoid [LinearOrderedCancelCommMonoid M] (S : section Preorder variable (M) -variable [Monoid M] [Preorder M] [CovariantClass M M (· * ·) (· ≤ ·)] {a : M} +variable [Monoid M] [Preorder M] [MulLeftMono M] {a : M} /-- The submonoid of elements greater than `1`. -/ @[to_additive (attr := simps) nonneg "The submonoid of nonnegative elements."] diff --git a/Mathlib/Algebra/Order/Monoid/Unbundled/Basic.lean b/Mathlib/Algebra/Order/Monoid/Unbundled/Basic.lean index ca4ff6cabcb05..10493ab993ec0 100644 --- a/Mathlib/Algebra/Order/Monoid/Unbundled/Basic.lean +++ b/Mathlib/Algebra/Order/Monoid/Unbundled/Basic.lean @@ -35,14 +35,14 @@ open Function section Nat -instance Nat.instCovariantClassMulLE : CovariantClass ℕ ℕ (· * ·) (· ≤ ·) where +instance Nat.instMulLeftMono : MulLeftMono ℕ where elim := fun _ _ _ h => mul_le_mul_left _ h end Nat section Int -instance Int.instCovariantClassAddLE : CovariantClass ℤ ℤ ((· + ·)) (· ≤ ·) where +instance Int.instAddLeftMono : AddLeftMono ℤ where elim := fun _ _ _ h => Int.add_le_add_left h _ end Int @@ -60,12 +60,12 @@ variable [LE α] /- The prime on this lemma is present only on the multiplicative version. The unprimed version is taken by the analogous lemma for semiring, with an extra non-negativity assumption. -/ @[to_additive (attr := gcongr) add_le_add_left] -theorem mul_le_mul_left' [CovariantClass α α (· * ·) (· ≤ ·)] {b c : α} (bc : b ≤ c) (a : α) : +theorem mul_le_mul_left' [MulLeftMono α] {b c : α} (bc : b ≤ c) (a : α) : a * b ≤ a * c := CovariantClass.elim _ bc @[to_additive le_of_add_le_add_left] -theorem le_of_mul_le_mul_left' [ContravariantClass α α (· * ·) (· ≤ ·)] {a b c : α} +theorem le_of_mul_le_mul_left' [MulLeftReflectLE α] {a b c : α} (bc : a * b ≤ a * c) : b ≤ c := ContravariantClass.elim _ bc @@ -73,26 +73,26 @@ theorem le_of_mul_le_mul_left' [ContravariantClass α α (· * ·) (· ≤ ·)] /- The prime on this lemma is present only on the multiplicative version. The unprimed version is taken by the analogous lemma for semiring, with an extra non-negativity assumption. -/ @[to_additive (attr := gcongr) add_le_add_right] -theorem mul_le_mul_right' [i : CovariantClass α α (swap (· * ·)) (· ≤ ·)] {b c : α} (bc : b ≤ c) +theorem mul_le_mul_right' [i : MulRightMono α] {b c : α} (bc : b ≤ c) (a : α) : b * a ≤ c * a := i.elim a bc @[to_additive le_of_add_le_add_right] -theorem le_of_mul_le_mul_right' [i : ContravariantClass α α (swap (· * ·)) (· ≤ ·)] {a b c : α} +theorem le_of_mul_le_mul_right' [i : MulRightReflectLE α] {a b c : α} (bc : b * a ≤ c * a) : b ≤ c := i.elim a bc @[to_additive (attr := simp)] -theorem mul_le_mul_iff_left [CovariantClass α α (· * ·) (· ≤ ·)] - [ContravariantClass α α (· * ·) (· ≤ ·)] (a : α) {b c : α} : +theorem mul_le_mul_iff_left [MulLeftMono α] + [MulLeftReflectLE α] (a : α) {b c : α} : a * b ≤ a * c ↔ b ≤ c := rel_iff_cov α α (· * ·) (· ≤ ·) a @[to_additive (attr := simp)] -theorem mul_le_mul_iff_right [CovariantClass α α (swap (· * ·)) (· ≤ ·)] - [ContravariantClass α α (swap (· * ·)) (· ≤ ·)] (a : α) {b c : α} : +theorem mul_le_mul_iff_right [MulRightMono α] + [MulRightReflectLE α] (a : α) {b c : α} : b * a ≤ c * a ↔ b ≤ c := rel_iff_cov α α (swap (· * ·)) (· ≤ ·) a @@ -103,36 +103,36 @@ section LT variable [LT α] @[to_additive (attr := simp)] -theorem mul_lt_mul_iff_left [CovariantClass α α (· * ·) (· < ·)] - [ContravariantClass α α (· * ·) (· < ·)] (a : α) {b c : α} : +theorem mul_lt_mul_iff_left [MulLeftStrictMono α] + [MulLeftReflectLT α] (a : α) {b c : α} : a * b < a * c ↔ b < c := rel_iff_cov α α (· * ·) (· < ·) a @[to_additive (attr := simp)] -theorem mul_lt_mul_iff_right [CovariantClass α α (swap (· * ·)) (· < ·)] - [ContravariantClass α α (swap (· * ·)) (· < ·)] (a : α) {b c : α} : +theorem mul_lt_mul_iff_right [MulRightStrictMono α] + [MulRightReflectLT α] (a : α) {b c : α} : b * a < c * a ↔ b < c := rel_iff_cov α α (swap (· * ·)) (· < ·) a @[to_additive (attr := gcongr) add_lt_add_left] -theorem mul_lt_mul_left' [CovariantClass α α (· * ·) (· < ·)] {b c : α} (bc : b < c) (a : α) : +theorem mul_lt_mul_left' [MulLeftStrictMono α] {b c : α} (bc : b < c) (a : α) : a * b < a * c := CovariantClass.elim _ bc @[to_additive lt_of_add_lt_add_left] -theorem lt_of_mul_lt_mul_left' [ContravariantClass α α (· * ·) (· < ·)] {a b c : α} +theorem lt_of_mul_lt_mul_left' [MulLeftReflectLT α] {a b c : α} (bc : a * b < a * c) : b < c := ContravariantClass.elim _ bc @[to_additive (attr := gcongr) add_lt_add_right] -theorem mul_lt_mul_right' [i : CovariantClass α α (swap (· * ·)) (· < ·)] {b c : α} (bc : b < c) +theorem mul_lt_mul_right' [i : MulRightStrictMono α] {b c : α} (bc : b < c) (a : α) : b * a < c * a := i.elim a bc @[to_additive lt_of_add_lt_add_right] -theorem lt_of_mul_lt_mul_right' [i : ContravariantClass α α (swap (· * ·)) (· < ·)] {a b c : α} +theorem lt_of_mul_lt_mul_right' [i : MulRightReflectLT α] {a b c : α} (bc : b * a < c * a) : b < c := i.elim a bc @@ -144,8 +144,8 @@ section Preorder variable [Preorder α] @[to_additive (attr := gcongr)] -theorem mul_lt_mul_of_lt_of_lt [CovariantClass α α (· * ·) (· < ·)] - [CovariantClass α α (swap (· * ·)) (· < ·)] +theorem mul_lt_mul_of_lt_of_lt [MulLeftStrictMono α] + [MulRightStrictMono α] {a b c d : α} (h₁ : a < b) (h₂ : c < d) : a * c < b * d := calc a * c < a * d := mul_lt_mul_left' h₂ a @@ -154,89 +154,89 @@ theorem mul_lt_mul_of_lt_of_lt [CovariantClass α α (· * ·) (· < ·)] alias add_lt_add := add_lt_add_of_lt_of_lt @[to_additive] -theorem mul_lt_mul_of_le_of_lt [CovariantClass α α (· * ·) (· < ·)] - [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a b c d : α} (h₁ : a ≤ b) (h₂ : c < d) : +theorem mul_lt_mul_of_le_of_lt [MulLeftStrictMono α] + [MulRightMono α] {a b c d : α} (h₁ : a ≤ b) (h₂ : c < d) : a * c < b * d := (mul_le_mul_right' h₁ _).trans_lt (mul_lt_mul_left' h₂ b) @[to_additive] -theorem mul_lt_mul_of_lt_of_le [CovariantClass α α (· * ·) (· ≤ ·)] - [CovariantClass α α (swap (· * ·)) (· < ·)] {a b c d : α} (h₁ : a < b) (h₂ : c ≤ d) : +theorem mul_lt_mul_of_lt_of_le [MulLeftMono α] + [MulRightStrictMono α] {a b c d : α} (h₁ : a < b) (h₂ : c ≤ d) : a * c < b * d := (mul_le_mul_left' h₂ _).trans_lt (mul_lt_mul_right' h₁ d) /-- Only assumes left strict covariance. -/ @[to_additive "Only assumes left strict covariance"] -theorem Left.mul_lt_mul [CovariantClass α α (· * ·) (· < ·)] - [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a b c d : α} (h₁ : a < b) (h₂ : c < d) : +theorem Left.mul_lt_mul [MulLeftStrictMono α] + [MulRightMono α] {a b c d : α} (h₁ : a < b) (h₂ : c < d) : a * c < b * d := mul_lt_mul_of_le_of_lt h₁.le h₂ /-- Only assumes right strict covariance. -/ @[to_additive "Only assumes right strict covariance"] -theorem Right.mul_lt_mul [CovariantClass α α (· * ·) (· ≤ ·)] - [CovariantClass α α (swap (· * ·)) (· < ·)] {a b c d : α} +theorem Right.mul_lt_mul [MulLeftMono α] + [MulRightStrictMono α] {a b c d : α} (h₁ : a < b) (h₂ : c < d) : a * c < b * d := mul_lt_mul_of_lt_of_le h₁ h₂.le @[to_additive (attr := gcongr) add_le_add] -theorem mul_le_mul' [CovariantClass α α (· * ·) (· ≤ ·)] [CovariantClass α α (swap (· * ·)) (· ≤ ·)] +theorem mul_le_mul' [MulLeftMono α] [MulRightMono α] {a b c d : α} (h₁ : a ≤ b) (h₂ : c ≤ d) : a * c ≤ b * d := (mul_le_mul_left' h₂ _).trans (mul_le_mul_right' h₁ d) @[to_additive] -theorem mul_le_mul_three [CovariantClass α α (· * ·) (· ≤ ·)] - [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a b c d e f : α} (h₁ : a ≤ d) (h₂ : b ≤ e) +theorem mul_le_mul_three [MulLeftMono α] + [MulRightMono α] {a b c d e f : α} (h₁ : a ≤ d) (h₂ : b ≤ e) (h₃ : c ≤ f) : a * b * c ≤ d * e * f := mul_le_mul' (mul_le_mul' h₁ h₂) h₃ @[to_additive] -theorem mul_lt_of_mul_lt_left [CovariantClass α α (· * ·) (· ≤ ·)] {a b c d : α} (h : a * b < c) +theorem mul_lt_of_mul_lt_left [MulLeftMono α] {a b c d : α} (h : a * b < c) (hle : d ≤ b) : a * d < c := (mul_le_mul_left' hle a).trans_lt h @[to_additive] -theorem mul_le_of_mul_le_left [CovariantClass α α (· * ·) (· ≤ ·)] {a b c d : α} (h : a * b ≤ c) +theorem mul_le_of_mul_le_left [MulLeftMono α] {a b c d : α} (h : a * b ≤ c) (hle : d ≤ b) : a * d ≤ c := @act_rel_of_rel_of_act_rel _ _ _ (· ≤ ·) _ _ a _ _ _ hle h @[to_additive] -theorem mul_lt_of_mul_lt_right [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a b c d : α} +theorem mul_lt_of_mul_lt_right [MulRightMono α] {a b c d : α} (h : a * b < c) (hle : d ≤ a) : d * b < c := (mul_le_mul_right' hle b).trans_lt h @[to_additive] -theorem mul_le_of_mul_le_right [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a b c d : α} +theorem mul_le_of_mul_le_right [MulRightMono α] {a b c d : α} (h : a * b ≤ c) (hle : d ≤ a) : d * b ≤ c := (mul_le_mul_right' hle b).trans h @[to_additive] -theorem lt_mul_of_lt_mul_left [CovariantClass α α (· * ·) (· ≤ ·)] {a b c d : α} (h : a < b * c) +theorem lt_mul_of_lt_mul_left [MulLeftMono α] {a b c d : α} (h : a < b * c) (hle : c ≤ d) : a < b * d := h.trans_le (mul_le_mul_left' hle b) @[to_additive] -theorem le_mul_of_le_mul_left [CovariantClass α α (· * ·) (· ≤ ·)] {a b c d : α} (h : a ≤ b * c) +theorem le_mul_of_le_mul_left [MulLeftMono α] {a b c d : α} (h : a ≤ b * c) (hle : c ≤ d) : a ≤ b * d := @rel_act_of_rel_of_rel_act _ _ _ (· ≤ ·) _ _ b _ _ _ hle h @[to_additive] -theorem lt_mul_of_lt_mul_right [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a b c d : α} +theorem lt_mul_of_lt_mul_right [MulRightMono α] {a b c d : α} (h : a < b * c) (hle : b ≤ d) : a < d * c := h.trans_le (mul_le_mul_right' hle c) @[to_additive] -theorem le_mul_of_le_mul_right [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a b c d : α} +theorem le_mul_of_le_mul_right [MulRightMono α] {a b c d : α} (h : a ≤ b * c) (hle : b ≤ d) : a ≤ d * c := h.trans (mul_le_mul_right' hle c) @@ -248,31 +248,31 @@ section PartialOrder variable [PartialOrder α] @[to_additive] -theorem mul_left_cancel'' [ContravariantClass α α (· * ·) (· ≤ ·)] {a b c : α} (h : a * b = a * c) : +theorem mul_left_cancel'' [MulLeftReflectLE α] {a b c : α} (h : a * b = a * c) : b = c := (le_of_mul_le_mul_left' h.le).antisymm (le_of_mul_le_mul_left' h.ge) @[to_additive] -theorem mul_right_cancel'' [ContravariantClass α α (swap (· * ·)) (· ≤ ·)] {a b c : α} +theorem mul_right_cancel'' [MulRightReflectLE α] {a b c : α} (h : a * b = c * b) : a = c := (le_of_mul_le_mul_right' h.le).antisymm (le_of_mul_le_mul_right' h.ge) -@[to_additive] lemma mul_le_mul_iff_of_ge [CovariantClass α α (· * ·) (· < ·)] - [CovariantClass α α (swap (· * ·)) (· < ·)] {a₁ a₂ b₁ b₂ : α} (ha : a₁ ≤ a₂) (hb : b₁ ≤ b₂) : +@[to_additive] lemma mul_le_mul_iff_of_ge [MulLeftStrictMono α] + [MulRightStrictMono α] {a₁ a₂ b₁ b₂ : α} (ha : a₁ ≤ a₂) (hb : b₁ ≤ b₂) : a₂ * b₂ ≤ a₁ * b₁ ↔ a₁ = a₂ ∧ b₁ = b₂ := by - haveI := covariantClass_le_of_lt α α (· * ·) - haveI := covariantClass_le_of_lt α α (swap (· * ·)) + haveI := mulLeftMono_of_mulLeftStrictMono α + haveI := mulRightMono_of_mulRightStrictMono α refine ⟨fun h ↦ ?_, by rintro ⟨rfl, rfl⟩; rfl⟩ simp only [eq_iff_le_not_lt, ha, hb, true_and] refine ⟨fun ha ↦ h.not_lt ?_, fun hb ↦ h.not_lt ?_⟩ exacts [mul_lt_mul_of_lt_of_le ha hb, mul_lt_mul_of_le_of_lt ha hb] -@[to_additive] theorem mul_eq_mul_iff_eq_and_eq [CovariantClass α α (· * ·) (· < ·)] - [CovariantClass α α (swap (· * ·)) (· < ·)] {a b c d : α} (hac : a ≤ c) (hbd : b ≤ d) : +@[to_additive] theorem mul_eq_mul_iff_eq_and_eq [MulLeftStrictMono α] + [MulRightStrictMono α] {a b c d : α} (hac : a ≤ c) (hbd : b ≤ d) : a * b = c * d ↔ a = c ∧ b = d := by - haveI := covariantClass_le_of_lt α α (· * ·) - haveI := covariantClass_le_of_lt α α (swap (· * ·)) + haveI := mulLeftMono_of_mulLeftStrictMono α + haveI := mulRightMono_of_mulRightStrictMono α rw [le_antisymm_iff, eq_true (mul_le_mul' hac hbd), true_and, mul_le_mul_iff_of_ge hac hbd] end PartialOrder @@ -281,31 +281,30 @@ section LinearOrder variable [LinearOrder α] {a b c d : α} @[to_additive] lemma min_lt_max_of_mul_lt_mul - [CovariantClass α α (· * ·) (· ≤ ·)] [CovariantClass α α (swap (· * ·)) (· ≤ ·)] + [MulLeftMono α] [MulRightMono α] (h : a * b < c * d) : min a b < max c d := by simp_rw [min_lt_iff, lt_max_iff]; contrapose! h; exact mul_le_mul' h.1.1 h.2.2 @[to_additive] lemma Left.min_le_max_of_mul_le_mul - [CovariantClass α α (· * ·) (· < ·)] [CovariantClass α α (swap (· * ·)) (· ≤ ·)] + [MulLeftStrictMono α] [MulRightMono α] (h : a * b ≤ c * d) : min a b ≤ max c d := by simp_rw [min_le_iff, le_max_iff]; contrapose! h; exact mul_lt_mul_of_le_of_lt h.1.1.le h.2.2 @[to_additive] lemma Right.min_le_max_of_mul_le_mul - [CovariantClass α α (· * ·) (· ≤ ·)] [CovariantClass α α (swap (· * ·)) (· < ·)] + [MulLeftMono α] [MulRightStrictMono α] (h : a * b ≤ c * d) : min a b ≤ max c d := by simp_rw [min_le_iff, le_max_iff]; contrapose! h; exact mul_lt_mul_of_lt_of_le h.1.1 h.2.2.le @[to_additive] lemma min_le_max_of_mul_le_mul - [CovariantClass α α (· * ·) (· < ·)] [CovariantClass α α (swap (· * ·)) (· < ·)] + [MulLeftStrictMono α] [MulRightStrictMono α] (h : a * b ≤ c * d) : min a b ≤ max c d := - let _ := covariantClass_le_of_lt α α (swap (· * ·)) + haveI := mulRightMono_of_mulRightStrictMono α Left.min_le_max_of_mul_le_mul h end LinearOrder section LinearOrder -variable [LinearOrder α] [CovariantClass α α (· * ·) (· ≤ ·)] - [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a b c d : α} +variable [LinearOrder α] [MulLeftMono α] [MulRightMono α] {a b c d : α} @[to_additive max_add_add_le_max_add_max] theorem max_mul_mul_le_max_mul_max' : max (a * b) (c * d) ≤ max a c * max b d := @@ -330,76 +329,76 @@ section LE variable [LE α] @[to_additive le_add_of_nonneg_right] -theorem le_mul_of_one_le_right' [CovariantClass α α (· * ·) (· ≤ ·)] {a b : α} (h : 1 ≤ b) : +theorem le_mul_of_one_le_right' [MulLeftMono α] {a b : α} (h : 1 ≤ b) : a ≤ a * b := calc a = a * 1 := (mul_one a).symm _ ≤ a * b := mul_le_mul_left' h a @[to_additive add_le_of_nonpos_right] -theorem mul_le_of_le_one_right' [CovariantClass α α (· * ·) (· ≤ ·)] {a b : α} (h : b ≤ 1) : +theorem mul_le_of_le_one_right' [MulLeftMono α] {a b : α} (h : b ≤ 1) : a * b ≤ a := calc a * b ≤ a * 1 := mul_le_mul_left' h a _ = a := mul_one a @[to_additive le_add_of_nonneg_left] -theorem le_mul_of_one_le_left' [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a b : α} (h : 1 ≤ b) : +theorem le_mul_of_one_le_left' [MulRightMono α] {a b : α} (h : 1 ≤ b) : a ≤ b * a := calc a = 1 * a := (one_mul a).symm _ ≤ b * a := mul_le_mul_right' h a @[to_additive add_le_of_nonpos_left] -theorem mul_le_of_le_one_left' [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a b : α} (h : b ≤ 1) : +theorem mul_le_of_le_one_left' [MulRightMono α] {a b : α} (h : b ≤ 1) : b * a ≤ a := calc b * a ≤ 1 * a := mul_le_mul_right' h a _ = a := one_mul a @[to_additive] -theorem one_le_of_le_mul_right [ContravariantClass α α (· * ·) (· ≤ ·)] {a b : α} (h : a ≤ a * b) : +theorem one_le_of_le_mul_right [MulLeftReflectLE α] {a b : α} (h : a ≤ a * b) : 1 ≤ b := le_of_mul_le_mul_left' (a := a) <| by simpa only [mul_one] @[to_additive] -theorem le_one_of_mul_le_right [ContravariantClass α α (· * ·) (· ≤ ·)] {a b : α} (h : a * b ≤ a) : +theorem le_one_of_mul_le_right [MulLeftReflectLE α] {a b : α} (h : a * b ≤ a) : b ≤ 1 := le_of_mul_le_mul_left' (a := a) <| by simpa only [mul_one] @[to_additive] -theorem one_le_of_le_mul_left [ContravariantClass α α (swap (· * ·)) (· ≤ ·)] {a b : α} +theorem one_le_of_le_mul_left [MulRightReflectLE α] {a b : α} (h : b ≤ a * b) : 1 ≤ a := le_of_mul_le_mul_right' (a := b) <| by simpa only [one_mul] @[to_additive] -theorem le_one_of_mul_le_left [ContravariantClass α α (swap (· * ·)) (· ≤ ·)] {a b : α} +theorem le_one_of_mul_le_left [MulRightReflectLE α] {a b : α} (h : a * b ≤ b) : a ≤ 1 := le_of_mul_le_mul_right' (a := b) <| by simpa only [one_mul] @[to_additive (attr := simp) le_add_iff_nonneg_right] -theorem le_mul_iff_one_le_right' [CovariantClass α α (· * ·) (· ≤ ·)] - [ContravariantClass α α (· * ·) (· ≤ ·)] (a : α) {b : α} : +theorem le_mul_iff_one_le_right' [MulLeftMono α] + [MulLeftReflectLE α] (a : α) {b : α} : a ≤ a * b ↔ 1 ≤ b := Iff.trans (by rw [mul_one]) (mul_le_mul_iff_left a) @[to_additive (attr := simp) le_add_iff_nonneg_left] -theorem le_mul_iff_one_le_left' [CovariantClass α α (swap (· * ·)) (· ≤ ·)] - [ContravariantClass α α (swap (· * ·)) (· ≤ ·)] (a : α) {b : α} : +theorem le_mul_iff_one_le_left' [MulRightMono α] + [MulRightReflectLE α] (a : α) {b : α} : a ≤ b * a ↔ 1 ≤ b := Iff.trans (by rw [one_mul]) (mul_le_mul_iff_right a) @[to_additive (attr := simp) add_le_iff_nonpos_right] -theorem mul_le_iff_le_one_right' [CovariantClass α α (· * ·) (· ≤ ·)] - [ContravariantClass α α (· * ·) (· ≤ ·)] (a : α) {b : α} : +theorem mul_le_iff_le_one_right' [MulLeftMono α] + [MulLeftReflectLE α] (a : α) {b : α} : a * b ≤ a ↔ b ≤ 1 := Iff.trans (by rw [mul_one]) (mul_le_mul_iff_left a) @[to_additive (attr := simp) add_le_iff_nonpos_left] -theorem mul_le_iff_le_one_left' [CovariantClass α α (swap (· * ·)) (· ≤ ·)] - [ContravariantClass α α (swap (· * ·)) (· ≤ ·)] {a b : α} : +theorem mul_le_iff_le_one_left' [MulRightMono α] + [MulRightReflectLE α] {a b : α} : a * b ≤ b ↔ a ≤ 1 := Iff.trans (by rw [one_mul]) (mul_le_mul_iff_right b) @@ -410,21 +409,21 @@ section LT variable [LT α] @[to_additive lt_add_of_pos_right] -theorem lt_mul_of_one_lt_right' [CovariantClass α α (· * ·) (· < ·)] (a : α) {b : α} (h : 1 < b) : +theorem lt_mul_of_one_lt_right' [MulLeftStrictMono α] (a : α) {b : α} (h : 1 < b) : a < a * b := calc a = a * 1 := (mul_one a).symm _ < a * b := mul_lt_mul_left' h a @[to_additive add_lt_of_neg_right] -theorem mul_lt_of_lt_one_right' [CovariantClass α α (· * ·) (· < ·)] (a : α) {b : α} (h : b < 1) : +theorem mul_lt_of_lt_one_right' [MulLeftStrictMono α] (a : α) {b : α} (h : b < 1) : a * b < a := calc a * b < a * 1 := mul_lt_mul_left' h a _ = a := mul_one a @[to_additive lt_add_of_pos_left] -theorem lt_mul_of_one_lt_left' [CovariantClass α α (swap (· * ·)) (· < ·)] (a : α) {b : α} +theorem lt_mul_of_one_lt_left' [MulRightStrictMono α] (a : α) {b : α} (h : 1 < b) : a < b * a := calc @@ -432,7 +431,7 @@ theorem lt_mul_of_one_lt_left' [CovariantClass α α (swap (· * ·)) (· < ·)] _ < b * a := mul_lt_mul_right' h a @[to_additive add_lt_of_neg_left] -theorem mul_lt_of_lt_one_left' [CovariantClass α α (swap (· * ·)) (· < ·)] (a : α) {b : α} +theorem mul_lt_of_lt_one_left' [MulRightStrictMono α] (a : α) {b : α} (h : b < 1) : b * a < a := calc @@ -440,47 +439,47 @@ theorem mul_lt_of_lt_one_left' [CovariantClass α α (swap (· * ·)) (· < ·)] _ = a := one_mul a @[to_additive] -theorem one_lt_of_lt_mul_right [ContravariantClass α α (· * ·) (· < ·)] {a b : α} (h : a < a * b) : +theorem one_lt_of_lt_mul_right [MulLeftReflectLT α] {a b : α} (h : a < a * b) : 1 < b := lt_of_mul_lt_mul_left' (a := a) <| by simpa only [mul_one] @[to_additive] -theorem lt_one_of_mul_lt_right [ContravariantClass α α (· * ·) (· < ·)] {a b : α} (h : a * b < a) : +theorem lt_one_of_mul_lt_right [MulLeftReflectLT α] {a b : α} (h : a * b < a) : b < 1 := lt_of_mul_lt_mul_left' (a := a) <| by simpa only [mul_one] @[to_additive] -theorem one_lt_of_lt_mul_left [ContravariantClass α α (swap (· * ·)) (· < ·)] {a b : α} +theorem one_lt_of_lt_mul_left [MulRightReflectLT α] {a b : α} (h : b < a * b) : 1 < a := lt_of_mul_lt_mul_right' (a := b) <| by simpa only [one_mul] @[to_additive] -theorem lt_one_of_mul_lt_left [ContravariantClass α α (swap (· * ·)) (· < ·)] {a b : α} +theorem lt_one_of_mul_lt_left [MulRightReflectLT α] {a b : α} (h : a * b < b) : a < 1 := lt_of_mul_lt_mul_right' (a := b) <| by simpa only [one_mul] @[to_additive (attr := simp) lt_add_iff_pos_right] -theorem lt_mul_iff_one_lt_right' [CovariantClass α α (· * ·) (· < ·)] - [ContravariantClass α α (· * ·) (· < ·)] (a : α) {b : α} : +theorem lt_mul_iff_one_lt_right' [MulLeftStrictMono α] + [MulLeftReflectLT α] (a : α) {b : α} : a < a * b ↔ 1 < b := Iff.trans (by rw [mul_one]) (mul_lt_mul_iff_left a) @[to_additive (attr := simp) lt_add_iff_pos_left] -theorem lt_mul_iff_one_lt_left' [CovariantClass α α (swap (· * ·)) (· < ·)] - [ContravariantClass α α (swap (· * ·)) (· < ·)] (a : α) {b : α} : a < b * a ↔ 1 < b := +theorem lt_mul_iff_one_lt_left' [MulRightStrictMono α] + [MulRightReflectLT α] (a : α) {b : α} : a < b * a ↔ 1 < b := Iff.trans (by rw [one_mul]) (mul_lt_mul_iff_right a) @[to_additive (attr := simp) add_lt_iff_neg_left] -theorem mul_lt_iff_lt_one_left' [CovariantClass α α (· * ·) (· < ·)] - [ContravariantClass α α (· * ·) (· < ·)] {a b : α} : +theorem mul_lt_iff_lt_one_left' [MulLeftStrictMono α] + [MulLeftReflectLT α] {a b : α} : a * b < a ↔ b < 1 := Iff.trans (by rw [mul_one]) (mul_lt_mul_iff_left a) @[to_additive (attr := simp) add_lt_iff_neg_right] -theorem mul_lt_iff_lt_one_right' [CovariantClass α α (swap (· * ·)) (· < ·)] - [ContravariantClass α α (swap (· * ·)) (· < ·)] {a : α} (b : α) : a * b < b ↔ a < 1 := +theorem mul_lt_iff_lt_one_right' [MulRightStrictMono α] + [MulRightReflectLT α] {a : α} (b : α) : a * b < b ↔ a < 1 := Iff.trans (by rw [one_mul]) (mul_lt_mul_iff_right b) end LT @@ -494,7 +493,7 @@ which assume left covariance. -/ @[to_additive] -theorem mul_le_of_le_of_le_one [CovariantClass α α (· * ·) (· ≤ ·)] {a b c : α} (hbc : b ≤ c) +theorem mul_le_of_le_of_le_one [MulLeftMono α] {a b c : α} (hbc : b ≤ c) (ha : a ≤ 1) : b * a ≤ c := calc @@ -503,7 +502,7 @@ theorem mul_le_of_le_of_le_one [CovariantClass α α (· * ·) (· ≤ ·)] {a b _ ≤ c := hbc @[to_additive] -theorem mul_lt_of_le_of_lt_one [CovariantClass α α (· * ·) (· < ·)] {a b c : α} (hbc : b ≤ c) +theorem mul_lt_of_le_of_lt_one [MulLeftStrictMono α] {a b c : α} (hbc : b ≤ c) (ha : a < 1) : b * a < c := calc @@ -512,7 +511,7 @@ theorem mul_lt_of_le_of_lt_one [CovariantClass α α (· * ·) (· < ·)] {a b c _ ≤ c := hbc @[to_additive] -theorem mul_lt_of_lt_of_le_one [CovariantClass α α (· * ·) (· ≤ ·)] {a b c : α} (hbc : b < c) +theorem mul_lt_of_lt_of_le_one [MulLeftMono α] {a b c : α} (hbc : b < c) (ha : a ≤ 1) : b * a < c := calc @@ -521,7 +520,7 @@ theorem mul_lt_of_lt_of_le_one [CovariantClass α α (· * ·) (· ≤ ·)] {a b _ < c := hbc @[to_additive] -theorem mul_lt_of_lt_of_lt_one [CovariantClass α α (· * ·) (· < ·)] {a b c : α} (hbc : b < c) +theorem mul_lt_of_lt_of_lt_one [MulLeftStrictMono α] {a b c : α} (hbc : b < c) (ha : a < 1) : b * a < c := calc @@ -530,7 +529,7 @@ theorem mul_lt_of_lt_of_lt_one [CovariantClass α α (· * ·) (· < ·)] {a b c _ < c := hbc @[to_additive] -theorem mul_lt_of_lt_of_lt_one' [CovariantClass α α (· * ·) (· ≤ ·)] {a b c : α} (hbc : b < c) +theorem mul_lt_of_lt_of_lt_one' [MulLeftMono α] {a b c : α} (hbc : b < c) (ha : a < 1) : b * a < c := mul_lt_of_lt_of_le_one hbc ha.le @@ -539,7 +538,7 @@ theorem mul_lt_of_lt_of_lt_one' [CovariantClass α α (· * ·) (· ≤ ·)] {a The lemma assuming right covariance is `Right.mul_le_one`. -/ @[to_additive "Assumes left covariance. The lemma assuming right covariance is `Right.add_nonpos`."] -theorem Left.mul_le_one [CovariantClass α α (· * ·) (· ≤ ·)] {a b : α} (ha : a ≤ 1) (hb : b ≤ 1) : +theorem Left.mul_le_one [MulLeftMono α] {a b : α} (ha : a ≤ 1) (hb : b ≤ 1) : a * b ≤ 1 := mul_le_of_le_of_le_one ha hb @@ -548,7 +547,7 @@ The lemma assuming right covariance is `Right.mul_lt_one_of_le_of_lt`. -/ @[to_additive Left.add_neg_of_nonpos_of_neg "Assumes left covariance. The lemma assuming right covariance is `Right.add_neg_of_nonpos_of_neg`."] -theorem Left.mul_lt_one_of_le_of_lt [CovariantClass α α (· * ·) (· < ·)] {a b : α} (ha : a ≤ 1) +theorem Left.mul_lt_one_of_le_of_lt [MulLeftStrictMono α] {a b : α} (ha : a ≤ 1) (hb : b < 1) : a * b < 1 := mul_lt_of_le_of_lt_one ha hb @@ -558,7 +557,7 @@ The lemma assuming right covariance is `Right.mul_lt_one_of_lt_of_le`. -/ @[to_additive Left.add_neg_of_neg_of_nonpos "Assumes left covariance. The lemma assuming right covariance is `Right.add_neg_of_neg_of_nonpos`."] -theorem Left.mul_lt_one_of_lt_of_le [CovariantClass α α (· * ·) (· ≤ ·)] {a b : α} (ha : a < 1) +theorem Left.mul_lt_one_of_lt_of_le [MulLeftMono α] {a b : α} (ha : a < 1) (hb : b ≤ 1) : a * b < 1 := mul_lt_of_lt_of_le_one ha hb @@ -567,7 +566,7 @@ theorem Left.mul_lt_one_of_lt_of_le [CovariantClass α α (· * ·) (· ≤ ·)] The lemma assuming right covariance is `Right.mul_lt_one`. -/ @[to_additive "Assumes left covariance. The lemma assuming right covariance is `Right.add_neg`."] -theorem Left.mul_lt_one [CovariantClass α α (· * ·) (· < ·)] {a b : α} (ha : a < 1) (hb : b < 1) : +theorem Left.mul_lt_one [MulLeftStrictMono α] {a b : α} (ha : a < 1) (hb : b < 1) : a * b < 1 := mul_lt_of_lt_of_lt_one ha hb @@ -575,7 +574,7 @@ theorem Left.mul_lt_one [CovariantClass α α (· * ·) (· < ·)] {a b : α} (h The lemma assuming right covariance is `Right.mul_lt_one'`. -/ @[to_additive "Assumes left covariance. The lemma assuming right covariance is `Right.add_neg'`."] -theorem Left.mul_lt_one' [CovariantClass α α (· * ·) (· ≤ ·)] {a b : α} (ha : a < 1) (hb : b < 1) : +theorem Left.mul_lt_one' [MulLeftMono α] {a b : α} (ha : a < 1) (hb : b < 1) : a * b < 1 := mul_lt_of_lt_of_lt_one' ha hb @@ -584,7 +583,7 @@ which assume left covariance. -/ @[to_additive] -theorem le_mul_of_le_of_one_le [CovariantClass α α (· * ·) (· ≤ ·)] {a b c : α} (hbc : b ≤ c) +theorem le_mul_of_le_of_one_le [MulLeftMono α] {a b c : α} (hbc : b ≤ c) (ha : 1 ≤ a) : b ≤ c * a := calc @@ -593,7 +592,7 @@ theorem le_mul_of_le_of_one_le [CovariantClass α α (· * ·) (· ≤ ·)] {a b _ ≤ c * a := mul_le_mul_left' ha c @[to_additive] -theorem lt_mul_of_le_of_one_lt [CovariantClass α α (· * ·) (· < ·)] {a b c : α} (hbc : b ≤ c) +theorem lt_mul_of_le_of_one_lt [MulLeftStrictMono α] {a b c : α} (hbc : b ≤ c) (ha : 1 < a) : b < c * a := calc @@ -602,7 +601,7 @@ theorem lt_mul_of_le_of_one_lt [CovariantClass α α (· * ·) (· < ·)] {a b c _ < c * a := mul_lt_mul_left' ha c @[to_additive] -theorem lt_mul_of_lt_of_one_le [CovariantClass α α (· * ·) (· ≤ ·)] {a b c : α} (hbc : b < c) +theorem lt_mul_of_lt_of_one_le [MulLeftMono α] {a b c : α} (hbc : b < c) (ha : 1 ≤ a) : b < c * a := calc @@ -611,7 +610,7 @@ theorem lt_mul_of_lt_of_one_le [CovariantClass α α (· * ·) (· ≤ ·)] {a b _ ≤ c * a := mul_le_mul_left' ha c @[to_additive] -theorem lt_mul_of_lt_of_one_lt [CovariantClass α α (· * ·) (· < ·)] {a b c : α} (hbc : b < c) +theorem lt_mul_of_lt_of_one_lt [MulLeftStrictMono α] {a b c : α} (hbc : b < c) (ha : 1 < a) : b < c * a := calc @@ -620,7 +619,7 @@ theorem lt_mul_of_lt_of_one_lt [CovariantClass α α (· * ·) (· < ·)] {a b c _ < c * a := mul_lt_mul_left' ha c @[to_additive] -theorem lt_mul_of_lt_of_one_lt' [CovariantClass α α (· * ·) (· ≤ ·)] {a b c : α} (hbc : b < c) +theorem lt_mul_of_lt_of_one_lt' [MulLeftMono α] {a b c : α} (hbc : b < c) (ha : 1 < a) : b < c * a := lt_mul_of_lt_of_one_le hbc ha.le @@ -629,7 +628,7 @@ theorem lt_mul_of_lt_of_one_lt' [CovariantClass α α (· * ·) (· ≤ ·)] {a The lemma assuming right covariance is `Right.one_le_mul`. -/ @[to_additive Left.add_nonneg "Assumes left covariance. The lemma assuming right covariance is `Right.add_nonneg`."] -theorem Left.one_le_mul [CovariantClass α α (· * ·) (· ≤ ·)] {a b : α} (ha : 1 ≤ a) (hb : 1 ≤ b) : +theorem Left.one_le_mul [MulLeftMono α] {a b : α} (ha : 1 ≤ a) (hb : 1 ≤ b) : 1 ≤ a * b := le_mul_of_le_of_one_le ha hb @@ -638,7 +637,7 @@ The lemma assuming right covariance is `Right.one_lt_mul_of_le_of_lt`. -/ @[to_additive Left.add_pos_of_nonneg_of_pos "Assumes left covariance. The lemma assuming right covariance is `Right.add_pos_of_nonneg_of_pos`."] -theorem Left.one_lt_mul_of_le_of_lt [CovariantClass α α (· * ·) (· < ·)] {a b : α} (ha : 1 ≤ a) +theorem Left.one_lt_mul_of_le_of_lt [MulLeftStrictMono α] {a b : α} (ha : 1 ≤ a) (hb : 1 < b) : 1 < a * b := lt_mul_of_le_of_one_lt ha hb @@ -648,7 +647,7 @@ The lemma assuming right covariance is `Right.one_lt_mul_of_lt_of_le`. -/ @[to_additive Left.add_pos_of_pos_of_nonneg "Assumes left covariance. The lemma assuming right covariance is `Right.add_pos_of_pos_of_nonneg`."] -theorem Left.one_lt_mul_of_lt_of_le [CovariantClass α α (· * ·) (· ≤ ·)] {a b : α} (ha : 1 < a) +theorem Left.one_lt_mul_of_lt_of_le [MulLeftMono α] {a b : α} (ha : 1 < a) (hb : 1 ≤ b) : 1 < a * b := lt_mul_of_lt_of_one_le ha hb @@ -657,7 +656,7 @@ theorem Left.one_lt_mul_of_lt_of_le [CovariantClass α α (· * ·) (· ≤ ·)] The lemma assuming right covariance is `Right.one_lt_mul`. -/ @[to_additive Left.add_pos "Assumes left covariance. The lemma assuming right covariance is `Right.add_pos`."] -theorem Left.one_lt_mul [CovariantClass α α (· * ·) (· < ·)] {a b : α} (ha : 1 < a) (hb : 1 < b) : +theorem Left.one_lt_mul [MulLeftStrictMono α] {a b : α} (ha : 1 < a) (hb : 1 < b) : 1 < a * b := lt_mul_of_lt_of_one_lt ha hb @@ -665,7 +664,7 @@ theorem Left.one_lt_mul [CovariantClass α α (· * ·) (· < ·)] {a b : α} (h The lemma assuming right covariance is `Right.one_lt_mul'`. -/ @[to_additive Left.add_pos' "Assumes left covariance. The lemma assuming right covariance is `Right.add_pos'`."] -theorem Left.one_lt_mul' [CovariantClass α α (· * ·) (· ≤ ·)] {a b : α} (ha : 1 < a) (hb : 1 < b) : +theorem Left.one_lt_mul' [MulLeftMono α] {a b : α} (ha : 1 < a) (hb : 1 < b) : 1 < a * b := lt_mul_of_lt_of_one_lt' ha hb @@ -674,7 +673,7 @@ which assume right covariance. -/ @[to_additive] -theorem mul_le_of_le_one_of_le [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a b c : α} (ha : a ≤ 1) +theorem mul_le_of_le_one_of_le [MulRightMono α] {a b c : α} (ha : a ≤ 1) (hbc : b ≤ c) : a * b ≤ c := calc @@ -683,7 +682,7 @@ theorem mul_le_of_le_one_of_le [CovariantClass α α (swap (· * ·)) (· ≤ · _ ≤ c := hbc @[to_additive] -theorem mul_lt_of_lt_one_of_le [CovariantClass α α (swap (· * ·)) (· < ·)] {a b c : α} (ha : a < 1) +theorem mul_lt_of_lt_one_of_le [MulRightStrictMono α] {a b c : α} (ha : a < 1) (hbc : b ≤ c) : a * b < c := calc @@ -692,7 +691,7 @@ theorem mul_lt_of_lt_one_of_le [CovariantClass α α (swap (· * ·)) (· < ·)] _ ≤ c := hbc @[to_additive] -theorem mul_lt_of_le_one_of_lt [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a b c : α} (ha : a ≤ 1) +theorem mul_lt_of_le_one_of_lt [MulRightMono α] {a b c : α} (ha : a ≤ 1) (hb : b < c) : a * b < c := calc @@ -701,7 +700,7 @@ theorem mul_lt_of_le_one_of_lt [CovariantClass α α (swap (· * ·)) (· ≤ · _ < c := hb @[to_additive] -theorem mul_lt_of_lt_one_of_lt [CovariantClass α α (swap (· * ·)) (· < ·)] {a b c : α} (ha : a < 1) +theorem mul_lt_of_lt_one_of_lt [MulRightStrictMono α] {a b c : α} (ha : a < 1) (hb : b < c) : a * b < c := calc @@ -710,7 +709,7 @@ theorem mul_lt_of_lt_one_of_lt [CovariantClass α α (swap (· * ·)) (· < ·)] _ < c := hb @[to_additive] -theorem mul_lt_of_lt_one_of_lt' [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a b c : α} (ha : a < 1) +theorem mul_lt_of_lt_one_of_lt' [MulRightMono α] {a b c : α} (ha : a < 1) (hbc : b < c) : a * b < c := mul_lt_of_le_one_of_lt ha.le hbc @@ -719,7 +718,7 @@ theorem mul_lt_of_lt_one_of_lt' [CovariantClass α α (swap (· * ·)) (· ≤ The lemma assuming left covariance is `Left.mul_le_one`. -/ @[to_additive "Assumes right covariance. The lemma assuming left covariance is `Left.add_nonpos`."] -theorem Right.mul_le_one [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a b : α} (ha : a ≤ 1) +theorem Right.mul_le_one [MulRightMono α] {a b : α} (ha : a ≤ 1) (hb : b ≤ 1) : a * b ≤ 1 := mul_le_of_le_one_of_le ha hb @@ -729,7 +728,7 @@ The lemma assuming left covariance is `Left.mul_lt_one_of_lt_of_le`. -/ @[to_additive Right.add_neg_of_neg_of_nonpos "Assumes right covariance. The lemma assuming left covariance is `Left.add_neg_of_neg_of_nonpos`."] -theorem Right.mul_lt_one_of_lt_of_le [CovariantClass α α (swap (· * ·)) (· < ·)] {a b : α} +theorem Right.mul_lt_one_of_lt_of_le [MulRightStrictMono α] {a b : α} (ha : a < 1) (hb : b ≤ 1) : a * b < 1 := mul_lt_of_lt_one_of_le ha hb @@ -739,7 +738,7 @@ The lemma assuming left covariance is `Left.mul_lt_one_of_le_of_lt`. -/ @[to_additive Right.add_neg_of_nonpos_of_neg "Assumes right covariance. The lemma assuming left covariance is `Left.add_neg_of_nonpos_of_neg`."] -theorem Right.mul_lt_one_of_le_of_lt [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a b : α} +theorem Right.mul_lt_one_of_le_of_lt [MulRightMono α] {a b : α} (ha : a ≤ 1) (hb : b < 1) : a * b < 1 := mul_lt_of_le_one_of_lt ha hb @@ -748,7 +747,7 @@ theorem Right.mul_lt_one_of_le_of_lt [CovariantClass α α (swap (· * ·)) (· The lemma assuming left covariance is `Left.mul_lt_one`. -/ @[to_additive "Assumes right covariance. The lemma assuming left covariance is `Left.add_neg`."] -theorem Right.mul_lt_one [CovariantClass α α (swap (· * ·)) (· < ·)] {a b : α} (ha : a < 1) +theorem Right.mul_lt_one [MulRightStrictMono α] {a b : α} (ha : a < 1) (hb : b < 1) : a * b < 1 := mul_lt_of_lt_one_of_lt ha hb @@ -757,7 +756,7 @@ theorem Right.mul_lt_one [CovariantClass α α (swap (· * ·)) (· < ·)] {a b The lemma assuming left covariance is `Left.mul_lt_one'`. -/ @[to_additive "Assumes right covariance. The lemma assuming left covariance is `Left.add_neg'`."] -theorem Right.mul_lt_one' [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a b : α} (ha : a < 1) +theorem Right.mul_lt_one' [MulRightMono α] {a b : α} (ha : a < 1) (hb : b < 1) : a * b < 1 := mul_lt_of_lt_one_of_lt' ha hb @@ -767,7 +766,7 @@ which assume right covariance. -/ @[to_additive] -theorem le_mul_of_one_le_of_le [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a b c : α} (ha : 1 ≤ a) +theorem le_mul_of_one_le_of_le [MulRightMono α] {a b c : α} (ha : 1 ≤ a) (hbc : b ≤ c) : b ≤ a * c := calc @@ -776,7 +775,7 @@ theorem le_mul_of_one_le_of_le [CovariantClass α α (swap (· * ·)) (· ≤ · _ ≤ a * c := mul_le_mul_right' ha c @[to_additive] -theorem lt_mul_of_one_lt_of_le [CovariantClass α α (swap (· * ·)) (· < ·)] {a b c : α} (ha : 1 < a) +theorem lt_mul_of_one_lt_of_le [MulRightStrictMono α] {a b c : α} (ha : 1 < a) (hbc : b ≤ c) : b < a * c := calc @@ -785,7 +784,7 @@ theorem lt_mul_of_one_lt_of_le [CovariantClass α α (swap (· * ·)) (· < ·)] _ < a * c := mul_lt_mul_right' ha c @[to_additive] -theorem lt_mul_of_one_le_of_lt [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a b c : α} (ha : 1 ≤ a) +theorem lt_mul_of_one_le_of_lt [MulRightMono α] {a b c : α} (ha : 1 ≤ a) (hbc : b < c) : b < a * c := calc @@ -794,7 +793,7 @@ theorem lt_mul_of_one_le_of_lt [CovariantClass α α (swap (· * ·)) (· ≤ · _ ≤ a * c := mul_le_mul_right' ha c @[to_additive] -theorem lt_mul_of_one_lt_of_lt [CovariantClass α α (swap (· * ·)) (· < ·)] {a b c : α} (ha : 1 < a) +theorem lt_mul_of_one_lt_of_lt [MulRightStrictMono α] {a b c : α} (ha : 1 < a) (hbc : b < c) : b < a * c := calc @@ -803,7 +802,7 @@ theorem lt_mul_of_one_lt_of_lt [CovariantClass α α (swap (· * ·)) (· < ·)] _ < a * c := mul_lt_mul_right' ha c @[to_additive] -theorem lt_mul_of_one_lt_of_lt' [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a b c : α} (ha : 1 < a) +theorem lt_mul_of_one_lt_of_lt' [MulRightMono α] {a b c : α} (ha : 1 < a) (hbc : b < c) : b < a * c := lt_mul_of_one_le_of_lt ha.le hbc @@ -812,7 +811,7 @@ theorem lt_mul_of_one_lt_of_lt' [CovariantClass α α (swap (· * ·)) (· ≤ The lemma assuming left covariance is `Left.one_le_mul`. -/ @[to_additive Right.add_nonneg "Assumes right covariance. The lemma assuming left covariance is `Left.add_nonneg`."] -theorem Right.one_le_mul [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a b : α} (ha : 1 ≤ a) +theorem Right.one_le_mul [MulRightMono α] {a b : α} (ha : 1 ≤ a) (hb : 1 ≤ b) : 1 ≤ a * b := le_mul_of_one_le_of_le ha hb @@ -822,7 +821,7 @@ The lemma assuming left covariance is `Left.one_lt_mul_of_lt_of_le`. -/ @[to_additive Right.add_pos_of_pos_of_nonneg "Assumes right covariance. The lemma assuming left covariance is `Left.add_pos_of_pos_of_nonneg`."] -theorem Right.one_lt_mul_of_lt_of_le [CovariantClass α α (swap (· * ·)) (· < ·)] {a b : α} +theorem Right.one_lt_mul_of_lt_of_le [MulRightStrictMono α] {a b : α} (ha : 1 < a) (hb : 1 ≤ b) : 1 < a * b := lt_mul_of_one_lt_of_le ha hb @@ -832,7 +831,7 @@ The lemma assuming left covariance is `Left.one_lt_mul_of_le_of_lt`. -/ @[to_additive Right.add_pos_of_nonneg_of_pos "Assumes right covariance. The lemma assuming left covariance is `Left.add_pos_of_nonneg_of_pos`."] -theorem Right.one_lt_mul_of_le_of_lt [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a b : α} +theorem Right.one_lt_mul_of_le_of_lt [MulRightMono α] {a b : α} (ha : 1 ≤ a) (hb : 1 < b) : 1 < a * b := lt_mul_of_one_le_of_lt ha hb @@ -841,7 +840,7 @@ theorem Right.one_lt_mul_of_le_of_lt [CovariantClass α α (swap (· * ·)) (· The lemma assuming left covariance is `Left.one_lt_mul`. -/ @[to_additive Right.add_pos "Assumes right covariance. The lemma assuming left covariance is `Left.add_pos`."] -theorem Right.one_lt_mul [CovariantClass α α (swap (· * ·)) (· < ·)] {a b : α} (ha : 1 < a) +theorem Right.one_lt_mul [MulRightStrictMono α] {a b : α} (ha : 1 < a) (hb : 1 < b) : 1 < a * b := lt_mul_of_one_lt_of_lt ha hb @@ -850,7 +849,7 @@ theorem Right.one_lt_mul [CovariantClass α α (swap (· * ·)) (· < ·)] {a b The lemma assuming left covariance is `Left.one_lt_mul'`. -/ @[to_additive Right.add_pos' "Assumes right covariance. The lemma assuming left covariance is `Left.add_pos'`."] -theorem Right.one_lt_mul' [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a b : α} (ha : 1 < a) +theorem Right.one_lt_mul' [MulRightMono α] {a b : α} (ha : 1 < a) (hb : 1 < b) : 1 < a * b := lt_mul_of_one_lt_of_lt' ha hb @@ -900,49 +899,49 @@ attribute [to_additive add_pos "**Alias** of `Left.add_pos`."] one_lt_mul' attribute [to_additive add_pos' "**Alias** of `Left.add_pos'`."] one_lt_mul'' @[to_additive] -theorem lt_of_mul_lt_of_one_le_left [CovariantClass α α (· * ·) (· ≤ ·)] {a b c : α} (h : a * b < c) +theorem lt_of_mul_lt_of_one_le_left [MulLeftMono α] {a b c : α} (h : a * b < c) (hle : 1 ≤ b) : a < c := (le_mul_of_one_le_right' hle).trans_lt h @[to_additive] -theorem le_of_mul_le_of_one_le_left [CovariantClass α α (· * ·) (· ≤ ·)] {a b c : α} (h : a * b ≤ c) +theorem le_of_mul_le_of_one_le_left [MulLeftMono α] {a b c : α} (h : a * b ≤ c) (hle : 1 ≤ b) : a ≤ c := (le_mul_of_one_le_right' hle).trans h @[to_additive] -theorem lt_of_lt_mul_of_le_one_left [CovariantClass α α (· * ·) (· ≤ ·)] {a b c : α} (h : a < b * c) +theorem lt_of_lt_mul_of_le_one_left [MulLeftMono α] {a b c : α} (h : a < b * c) (hle : c ≤ 1) : a < b := h.trans_le (mul_le_of_le_one_right' hle) @[to_additive] -theorem le_of_le_mul_of_le_one_left [CovariantClass α α (· * ·) (· ≤ ·)] {a b c : α} (h : a ≤ b * c) +theorem le_of_le_mul_of_le_one_left [MulLeftMono α] {a b c : α} (h : a ≤ b * c) (hle : c ≤ 1) : a ≤ b := h.trans (mul_le_of_le_one_right' hle) @[to_additive] -theorem lt_of_mul_lt_of_one_le_right [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a b c : α} +theorem lt_of_mul_lt_of_one_le_right [MulRightMono α] {a b c : α} (h : a * b < c) (hle : 1 ≤ a) : b < c := (le_mul_of_one_le_left' hle).trans_lt h @[to_additive] -theorem le_of_mul_le_of_one_le_right [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a b c : α} +theorem le_of_mul_le_of_one_le_right [MulRightMono α] {a b c : α} (h : a * b ≤ c) (hle : 1 ≤ a) : b ≤ c := (le_mul_of_one_le_left' hle).trans h @[to_additive] -theorem lt_of_lt_mul_of_le_one_right [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a b c : α} +theorem lt_of_lt_mul_of_le_one_right [MulRightMono α] {a b c : α} (h : a < b * c) (hle : b ≤ 1) : a < c := h.trans_le (mul_le_of_le_one_left' hle) @[to_additive] -theorem le_of_le_mul_of_le_one_right [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a b c : α} +theorem le_of_le_mul_of_le_one_right [MulRightMono α] {a b c : α} (h : a ≤ b * c) (hle : b ≤ 1) : a ≤ c := h.trans (mul_le_of_le_one_left' hle) @@ -954,8 +953,8 @@ section PartialOrder variable [PartialOrder α] @[to_additive] -theorem mul_eq_one_iff_of_one_le [CovariantClass α α (· * ·) (· ≤ ·)] - [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a b : α} (ha : 1 ≤ a) (hb : 1 ≤ b) : +theorem mul_eq_one_iff_of_one_le [MulLeftMono α] + [MulRightMono α] {a b : α} (ha : 1 ≤ a) (hb : 1 ≤ b) : a * b = 1 ↔ a = 1 ∧ b = 1 := Iff.intro (fun hab : a * b = 1 => @@ -971,7 +970,7 @@ theorem mul_eq_one_iff_of_one_le [CovariantClass α α (· * ·) (· ≤ ·)] section Left -variable [CovariantClass α α (· * ·) (· ≤ ·)] {a b : α} +variable [MulLeftMono α] {a b : α} @[to_additive eq_zero_of_add_nonneg_left] theorem eq_one_of_one_le_mul_left (ha : a ≤ 1) (hb : b ≤ 1) (hab : 1 ≤ a * b) : a = 1 := @@ -985,7 +984,7 @@ end Left section Right -variable [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a b : α} +variable [MulRightMono α] {a b : α} @[to_additive eq_zero_of_add_nonneg_right] theorem eq_one_of_one_le_mul_right (ha : a ≤ 1) (hb : b ≤ 1) (hab : 1 ≤ a * b) : b = 1 := @@ -1003,7 +1002,7 @@ section LinearOrder variable [LinearOrder α] -theorem exists_square_le [CovariantClass α α (· * ·) (· < ·)] (a : α) : ∃ b : α, b * b ≤ a := by +theorem exists_square_le [MulLeftStrictMono α] (a : α) : ∃ b : α, b * b ≤ a := by by_cases h : a < 1 · use a have : a * a < a * 1 := mul_lt_mul_left' h a @@ -1026,24 +1025,24 @@ section PartialOrder variable [PartialOrder α] /- This is not instance, since we want to have an instance from `LeftCancelSemigroup`s -to the appropriate `CovariantClass`. -/ +to the appropriate covariant class. -/ /-- A semigroup with a partial order and satisfying `LeftCancelSemigroup` -(i.e. `a * c < b * c → a < b`) is a `left_cancel Semigroup`. -/ +(i.e. `a * c < b * c → a < b`) is a `LeftCancelSemigroup`. -/ @[to_additive "An additive semigroup with a partial order and satisfying `AddLeftCancelSemigroup` -(i.e. `c + a < c + b → a < b`) is a `left_cancel AddSemigroup`."] -def Contravariant.toLeftCancelSemigroup [ContravariantClass α α (· * ·) (· ≤ ·)] : +(i.e. `c + a < c + b → a < b`) is a `AddLeftCancelSemigroup`."] +def Contravariant.toLeftCancelSemigroup [MulLeftReflectLE α] : LeftCancelSemigroup α := { ‹Semigroup α› with mul_left_cancel := fun _ _ _ => mul_left_cancel'' } /- This is not instance, since we want to have an instance from `RightCancelSemigroup`s -to the appropriate `CovariantClass`. -/ +to the appropriate covariant class. -/ /-- A semigroup with a partial order and satisfying `RightCancelSemigroup` -(i.e. `a * c < b * c → a < b`) is a `right_cancel Semigroup`. -/ +(i.e. `a * c < b * c → a < b`) is a `RightCancelSemigroup`. -/ @[to_additive "An additive semigroup with a partial order and satisfying `AddRightCancelSemigroup` -(`a + c < b + c → a < b`) is a `right_cancel AddSemigroup`."] -def Contravariant.toRightCancelSemigroup [ContravariantClass α α (swap (· * ·)) (· ≤ ·)] : +(`a + c < b + c → a < b`) is a `AddRightCancelSemigroup`."] +def Contravariant.toRightCancelSemigroup [MulRightReflectLE α] : RightCancelSemigroup α := { ‹Semigroup α› with mul_right_cancel := fun _ _ _ => mul_right_cancel'' } @@ -1056,68 +1055,68 @@ section Mono variable [Mul α] [Preorder α] [Preorder β] {f g : β → α} {s : Set β} @[to_additive const_add] -theorem Monotone.const_mul' [CovariantClass α α (· * ·) (· ≤ ·)] (hf : Monotone f) (a : α) : +theorem Monotone.const_mul' [MulLeftMono α] (hf : Monotone f) (a : α) : Monotone fun x => a * f x := fun _ _ h => mul_le_mul_left' (hf h) a @[to_additive const_add] -theorem MonotoneOn.const_mul' [CovariantClass α α (· * ·) (· ≤ ·)] (hf : MonotoneOn f s) (a : α) : +theorem MonotoneOn.const_mul' [MulLeftMono α] (hf : MonotoneOn f s) (a : α) : MonotoneOn (fun x => a * f x) s := fun _ hx _ hy h => mul_le_mul_left' (hf hx hy h) a @[to_additive const_add] -theorem Antitone.const_mul' [CovariantClass α α (· * ·) (· ≤ ·)] (hf : Antitone f) (a : α) : +theorem Antitone.const_mul' [MulLeftMono α] (hf : Antitone f) (a : α) : Antitone fun x => a * f x := fun _ _ h => mul_le_mul_left' (hf h) a @[to_additive const_add] -theorem AntitoneOn.const_mul' [CovariantClass α α (· * ·) (· ≤ ·)] (hf : AntitoneOn f s) (a : α) : +theorem AntitoneOn.const_mul' [MulLeftMono α] (hf : AntitoneOn f s) (a : α) : AntitoneOn (fun x => a * f x) s := fun _ hx _ hy h => mul_le_mul_left' (hf hx hy h) a @[to_additive add_const] -theorem Monotone.mul_const' [CovariantClass α α (swap (· * ·)) (· ≤ ·)] (hf : Monotone f) (a : α) : +theorem Monotone.mul_const' [MulRightMono α] (hf : Monotone f) (a : α) : Monotone fun x => f x * a := fun _ _ h => mul_le_mul_right' (hf h) a @[to_additive add_const] -theorem MonotoneOn.mul_const' [CovariantClass α α (swap (· * ·)) (· ≤ ·)] (hf : MonotoneOn f s) +theorem MonotoneOn.mul_const' [MulRightMono α] (hf : MonotoneOn f s) (a : α) : MonotoneOn (fun x => f x * a) s := fun _ hx _ hy h => mul_le_mul_right' (hf hx hy h) a @[to_additive add_const] -theorem Antitone.mul_const' [CovariantClass α α (swap (· * ·)) (· ≤ ·)] (hf : Antitone f) (a : α) : +theorem Antitone.mul_const' [MulRightMono α] (hf : Antitone f) (a : α) : Antitone fun x => f x * a := fun _ _ h => mul_le_mul_right' (hf h) a @[to_additive add_const] -theorem AntitoneOn.mul_const' [CovariantClass α α (swap (· * ·)) (· ≤ ·)] (hf : AntitoneOn f s) +theorem AntitoneOn.mul_const' [MulRightMono α] (hf : AntitoneOn f s) (a : α) : AntitoneOn (fun x => f x * a) s := fun _ hx _ hy h => mul_le_mul_right' (hf hx hy h) a /-- The product of two monotone functions is monotone. -/ @[to_additive add "The sum of two monotone functions is monotone."] -theorem Monotone.mul' [CovariantClass α α (· * ·) (· ≤ ·)] - [CovariantClass α α (swap (· * ·)) (· ≤ ·)] (hf : Monotone f) (hg : Monotone g) : +theorem Monotone.mul' [MulLeftMono α] + [MulRightMono α] (hf : Monotone f) (hg : Monotone g) : Monotone fun x => f x * g x := fun _ _ h => mul_le_mul' (hf h) (hg h) /-- The product of two monotone functions is monotone. -/ @[to_additive add "The sum of two monotone functions is monotone."] -theorem MonotoneOn.mul' [CovariantClass α α (· * ·) (· ≤ ·)] - [CovariantClass α α (swap (· * ·)) (· ≤ ·)] (hf : MonotoneOn f s) (hg : MonotoneOn g s) : +theorem MonotoneOn.mul' [MulLeftMono α] + [MulRightMono α] (hf : MonotoneOn f s) (hg : MonotoneOn g s) : MonotoneOn (fun x => f x * g x) s := fun _ hx _ hy h => mul_le_mul' (hf hx hy h) (hg hx hy h) /-- The product of two antitone functions is antitone. -/ @[to_additive add "The sum of two antitone functions is antitone."] -theorem Antitone.mul' [CovariantClass α α (· * ·) (· ≤ ·)] - [CovariantClass α α (swap (· * ·)) (· ≤ ·)] (hf : Antitone f) (hg : Antitone g) : +theorem Antitone.mul' [MulLeftMono α] + [MulRightMono α] (hf : Antitone f) (hg : Antitone g) : Antitone fun x => f x * g x := fun _ _ h => mul_le_mul' (hf h) (hg h) /-- The product of two antitone functions is antitone. -/ @[to_additive add "The sum of two antitone functions is antitone."] -theorem AntitoneOn.mul' [CovariantClass α α (· * ·) (· ≤ ·)] - [CovariantClass α α (swap (· * ·)) (· ≤ ·)] (hf : AntitoneOn f s) (hg : AntitoneOn g s) : +theorem AntitoneOn.mul' [MulLeftMono α] + [MulRightMono α] (hf : AntitoneOn f s) (hg : AntitoneOn g s) : AntitoneOn (fun x => f x * g x) s := fun _ hx _ hy h => mul_le_mul' (hf hx hy h) (hg hx hy h) section Left -variable [CovariantClass α α (· * ·) (· < ·)] +variable [MulLeftStrictMono α] @[to_additive const_add] theorem StrictMono.const_mul' (hf : StrictMono f) (c : α) : StrictMono fun x => c * f x := @@ -1141,7 +1140,7 @@ end Left section Right -variable [CovariantClass α α (swap (· * ·)) (· < ·)] +variable [MulRightStrictMono α] @[to_additive add_const] theorem StrictMono.mul_const' (hf : StrictMono f) (c : α) : StrictMono fun x => f x * c := @@ -1165,37 +1164,37 @@ end Right /-- The product of two strictly monotone functions is strictly monotone. -/ @[to_additive add "The sum of two strictly monotone functions is strictly monotone."] -theorem StrictMono.mul' [CovariantClass α α (· * ·) (· < ·)] - [CovariantClass α α (swap (· * ·)) (· < ·)] (hf : StrictMono f) (hg : StrictMono g) : +theorem StrictMono.mul' [MulLeftStrictMono α] + [MulRightStrictMono α] (hf : StrictMono f) (hg : StrictMono g) : StrictMono fun x => f x * g x := fun _ _ ab => mul_lt_mul_of_lt_of_lt (hf ab) (hg ab) /-- The product of two strictly monotone functions is strictly monotone. -/ @[to_additive add "The sum of two strictly monotone functions is strictly monotone."] -theorem StrictMonoOn.mul' [CovariantClass α α (· * ·) (· < ·)] - [CovariantClass α α (swap (· * ·)) (· < ·)] (hf : StrictMonoOn f s) (hg : StrictMonoOn g s) : +theorem StrictMonoOn.mul' [MulLeftStrictMono α] + [MulRightStrictMono α] (hf : StrictMonoOn f s) (hg : StrictMonoOn g s) : StrictMonoOn (fun x => f x * g x) s := fun _ ha _ hb ab => mul_lt_mul_of_lt_of_lt (hf ha hb ab) (hg ha hb ab) /-- The product of two strictly antitone functions is strictly antitone. -/ @[to_additive add "The sum of two strictly antitone functions is strictly antitone."] -theorem StrictAnti.mul' [CovariantClass α α (· * ·) (· < ·)] - [CovariantClass α α (swap (· * ·)) (· < ·)] (hf : StrictAnti f) (hg : StrictAnti g) : +theorem StrictAnti.mul' [MulLeftStrictMono α] + [MulRightStrictMono α] (hf : StrictAnti f) (hg : StrictAnti g) : StrictAnti fun x => f x * g x := fun _ _ ab => mul_lt_mul_of_lt_of_lt (hf ab) (hg ab) /-- The product of two strictly antitone functions is strictly antitone. -/ @[to_additive add "The sum of two strictly antitone functions is strictly antitone."] -theorem StrictAntiOn.mul' [CovariantClass α α (· * ·) (· < ·)] - [CovariantClass α α (swap (· * ·)) (· < ·)] (hf : StrictAntiOn f s) (hg : StrictAntiOn g s) : +theorem StrictAntiOn.mul' [MulLeftStrictMono α] + [MulRightStrictMono α] (hf : StrictAntiOn f s) (hg : StrictAntiOn g s) : StrictAntiOn (fun x => f x * g x) s := fun _ ha _ hb ab => mul_lt_mul_of_lt_of_lt (hf ha hb ab) (hg ha hb ab) /-- The product of a monotone function and a strictly monotone function is strictly monotone. -/ @[to_additive add_strictMono "The sum of a monotone function and a strictly monotone function is strictly monotone."] -theorem Monotone.mul_strictMono' [CovariantClass α α (· * ·) (· < ·)] - [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {f g : β → α} (hf : Monotone f) +theorem Monotone.mul_strictMono' [MulLeftStrictMono α] + [MulRightMono α] {f g : β → α} (hf : Monotone f) (hg : StrictMono g) : StrictMono fun x => f x * g x := fun _ _ h => mul_lt_mul_of_le_of_lt (hf h.le) (hg h) @@ -1203,16 +1202,16 @@ theorem Monotone.mul_strictMono' [CovariantClass α α (· * ·) (· < ·)] /-- The product of a monotone function and a strictly monotone function is strictly monotone. -/ @[to_additive add_strictMono "The sum of a monotone function and a strictly monotone function is strictly monotone."] -theorem MonotoneOn.mul_strictMono' [CovariantClass α α (· * ·) (· < ·)] - [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {f g : β → α} (hf : MonotoneOn f s) +theorem MonotoneOn.mul_strictMono' [MulLeftStrictMono α] + [MulRightMono α] {f g : β → α} (hf : MonotoneOn f s) (hg : StrictMonoOn g s) : StrictMonoOn (fun x => f x * g x) s := fun _ hx _ hy h => mul_lt_mul_of_le_of_lt (hf hx hy h.le) (hg hx hy h) /-- The product of an antitone function and a strictly antitone function is strictly antitone. -/ @[to_additive add_strictAnti "The sum of an antitone function and a strictly antitone function is strictly antitone."] -theorem Antitone.mul_strictAnti' [CovariantClass α α (· * ·) (· < ·)] - [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {f g : β → α} (hf : Antitone f) +theorem Antitone.mul_strictAnti' [MulLeftStrictMono α] + [MulRightMono α] {f g : β → α} (hf : Antitone f) (hg : StrictAnti g) : StrictAnti fun x => f x * g x := fun _ _ h => mul_lt_mul_of_le_of_lt (hf h.le) (hg h) @@ -1220,13 +1219,13 @@ theorem Antitone.mul_strictAnti' [CovariantClass α α (· * ·) (· < ·)] /-- The product of an antitone function and a strictly antitone function is strictly antitone. -/ @[to_additive add_strictAnti "The sum of an antitone function and a strictly antitone function is strictly antitone."] -theorem AntitoneOn.mul_strictAnti' [CovariantClass α α (· * ·) (· < ·)] - [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {f g : β → α} (hf : AntitoneOn f s) +theorem AntitoneOn.mul_strictAnti' [MulLeftStrictMono α] + [MulRightMono α] {f g : β → α} (hf : AntitoneOn f s) (hg : StrictAntiOn g s) : StrictAntiOn (fun x => f x * g x) s := fun _ hx _ hy h => mul_lt_mul_of_le_of_lt (hf hx hy h.le) (hg hx hy h) -variable [CovariantClass α α (· * ·) (· ≤ ·)] [CovariantClass α α (swap (· * ·)) (· < ·)] +variable [MulLeftMono α] [MulRightStrictMono α] /-- The product of a strictly monotone function and a monotone function is strictly monotone. -/ @[to_additive add_monotone "The sum of a strictly monotone function and a monotone function is @@ -1257,34 +1256,34 @@ theorem StrictAntiOn.mul_antitone' (hf : StrictAntiOn f s) (hg : AntitoneOn g s) fun _ hx _ hy h => mul_lt_mul_of_lt_of_le (hf hx hy h) (hg hx hy h.le) @[to_additive (attr := simp) cmp_add_left] -theorem cmp_mul_left' {α : Type*} [Mul α] [LinearOrder α] [CovariantClass α α (· * ·) (· < ·)] +theorem cmp_mul_left' {α : Type*} [Mul α] [LinearOrder α] [MulLeftStrictMono α] (a b c : α) : cmp (a * b) (a * c) = cmp b c := (strictMono_id.const_mul' a).cmp_map_eq b c @[to_additive (attr := simp) cmp_add_right] theorem cmp_mul_right' {α : Type*} [Mul α] [LinearOrder α] - [CovariantClass α α (swap (· * ·)) (· < ·)] (a b c : α) : + [MulRightStrictMono α] (a b c : α) : cmp (a * c) (b * c) = cmp a b := (strictMono_id.mul_const' c).cmp_map_eq a b end Mono /-- An element `a : α` is `MulLECancellable` if `x ↦ a * x` is order-reflecting. -We will make a separate version of many lemmas that require `[ContravariantClass α α (*) (≤)]` with +We will make a separate version of many lemmas that require `[MulLeftReflectLE α]` with `MulLECancellable` assumptions instead. These lemmas can then be instantiated to specific types, like `ENNReal`, where we can replace the assumption `AddLECancellable x` by `x ≠ ∞`. -/ @[to_additive "An element `a : α` is `AddLECancellable` if `x ↦ a + x` is order-reflecting. -We will make a separate version of many lemmas that require `[ContravariantClass α α (+) (≤)]` with +We will make a separate version of many lemmas that require `[MulLeftReflectLE α]` with `AddLECancellable` assumptions instead. These lemmas can then be instantiated to specific types, like `ENNReal`, where we can replace the assumption `AddLECancellable x` by `x ≠ ∞`. "] def MulLECancellable [Mul α] [LE α] (a : α) : Prop := ∀ ⦃b c⦄, a * b ≤ a * c → b ≤ c @[to_additive] -theorem Contravariant.MulLECancellable [Mul α] [LE α] [ContravariantClass α α (· * ·) (· ≤ ·)] +theorem Contravariant.MulLECancellable [Mul α] [LE α] [MulLeftReflectLE α] {a : α} : MulLECancellable a := fun _ _ => le_of_mul_le_mul_left' @@ -1319,51 +1318,51 @@ protected theorem inj_left [Mul α] [@Std.Commutative α (· * ·)] [PartialOrde variable [LE α] @[to_additive] -protected theorem mul_le_mul_iff_left [Mul α] [CovariantClass α α (· * ·) (· ≤ ·)] {a b c : α} +protected theorem mul_le_mul_iff_left [Mul α] [MulLeftMono α] {a b c : α} (ha : MulLECancellable a) : a * b ≤ a * c ↔ b ≤ c := ⟨fun h => ha h, fun h => mul_le_mul_left' h a⟩ @[to_additive] protected theorem mul_le_mul_iff_right [Mul α] [i : @Std.Commutative α (· * ·)] - [CovariantClass α α (· * ·) (· ≤ ·)] {a b c : α} (ha : MulLECancellable a) : + [MulLeftMono α] {a b c : α} (ha : MulLECancellable a) : b * a ≤ c * a ↔ b ≤ c := by rw [i.comm b, i.comm c, ha.mul_le_mul_iff_left] @[to_additive] -protected theorem le_mul_iff_one_le_right [MulOneClass α] [CovariantClass α α (· * ·) (· ≤ ·)] +protected theorem le_mul_iff_one_le_right [MulOneClass α] [MulLeftMono α] {a b : α} (ha : MulLECancellable a) : a ≤ a * b ↔ 1 ≤ b := Iff.trans (by rw [mul_one]) ha.mul_le_mul_iff_left @[to_additive] -protected theorem mul_le_iff_le_one_right [MulOneClass α] [CovariantClass α α (· * ·) (· ≤ ·)] +protected theorem mul_le_iff_le_one_right [MulOneClass α] [MulLeftMono α] {a b : α} (ha : MulLECancellable a) : a * b ≤ a ↔ b ≤ 1 := Iff.trans (by rw [mul_one]) ha.mul_le_mul_iff_left @[to_additive] protected theorem le_mul_iff_one_le_left [MulOneClass α] [i : @Std.Commutative α (· * ·)] - [CovariantClass α α (· * ·) (· ≤ ·)] {a b : α} (ha : MulLECancellable a) : + [MulLeftMono α] {a b : α} (ha : MulLECancellable a) : a ≤ b * a ↔ 1 ≤ b := by rw [i.comm, ha.le_mul_iff_one_le_right] @[to_additive] protected theorem mul_le_iff_le_one_left [MulOneClass α] [i : @Std.Commutative α (· * ·)] - [CovariantClass α α (· * ·) (· ≤ ·)] {a b : α} (ha : MulLECancellable a) : + [MulLeftMono α] {a b : α} (ha : MulLECancellable a) : b * a ≤ a ↔ b ≤ 1 := by rw [i.comm, ha.mul_le_iff_le_one_right] @[to_additive] lemma mul [Semigroup α] {a b : α} (ha : MulLECancellable a) (hb : MulLECancellable b) : MulLECancellable (a * b) := fun c d hcd ↦ hb <| ha <| by rwa [← mul_assoc, ← mul_assoc] -@[to_additive] lemma of_mul_right [Semigroup α] [CovariantClass α α (· * ·) (· ≤ ·)] {a b : α} +@[to_additive] lemma of_mul_right [Semigroup α] [MulLeftMono α] {a b : α} (h : MulLECancellable (a * b)) : MulLECancellable b := fun c d hcd ↦ h <| by rw [mul_assoc, mul_assoc]; exact mul_le_mul_left' hcd _ -@[to_additive] lemma of_mul_left [CommSemigroup α] [CovariantClass α α (· * ·) (· ≤ ·)] {a b : α} +@[to_additive] lemma of_mul_left [CommSemigroup α] [MulLeftMono α] {a b : α} (h : MulLECancellable (a * b)) : MulLECancellable a := (mul_comm a b ▸ h).of_mul_right end MulLECancellable @[to_additive (attr := simp)] -lemma mulLECancellable_mul [LE α] [CommSemigroup α] [CovariantClass α α (· * ·) (· ≤ ·)] {a b : α} : +lemma mulLECancellable_mul [LE α] [CommSemigroup α] [MulLeftMono α] {a b : α} : MulLECancellable (a * b) ↔ MulLECancellable a ∧ MulLECancellable b := ⟨fun h ↦ ⟨h.of_mul_left, h.of_mul_right⟩, fun h ↦ h.1.mul h.2⟩ diff --git a/Mathlib/Algebra/Order/Monoid/Unbundled/Defs.lean b/Mathlib/Algebra/Order/Monoid/Unbundled/Defs.lean index 6418be782db5c..8c000a8d6ddd3 100644 --- a/Mathlib/Algebra/Order/Monoid/Unbundled/Defs.lean +++ b/Mathlib/Algebra/Order/Monoid/Unbundled/Defs.lean @@ -531,28 +531,28 @@ theorem contravariant_le_of_contravariant_eq_and_lt [PartialOrder N] then the following four instances (actually eight) can be removed in favor of the above two. -/ @[to_additive] -instance IsLeftCancelMul.covariant_mul_lt_of_covariant_mul_le [Mul N] [IsLeftCancelMul N] - [PartialOrder N] [CovariantClass N N (· * ·) (· ≤ ·)] : - CovariantClass N N (· * ·) (· < ·) where +instance IsLeftCancelMul.mulLeftStrictMono_of_mulLeftMono [Mul N] [IsLeftCancelMul N] + [PartialOrder N] [MulLeftMono N] : + MulLeftStrictMono N where elim a _ _ bc := (CovariantClass.elim a bc.le).lt_of_ne ((mul_ne_mul_right a).mpr bc.ne) @[to_additive] -instance IsRightCancelMul.covariant_swap_mul_lt_of_covariant_swap_mul_le - [Mul N] [IsRightCancelMul N] [PartialOrder N] [CovariantClass N N (swap (· * ·)) (· ≤ ·)] : - CovariantClass N N (swap (· * ·)) (· < ·) where +instance IsRightCancelMul.mulRightStrictMono_of_mulRightMono + [Mul N] [IsRightCancelMul N] [PartialOrder N] [MulRightMono N] : + MulRightStrictMono N where elim a _ _ bc := (CovariantClass.elim a bc.le).lt_of_ne ((mul_ne_mul_left a).mpr bc.ne) @[to_additive] -instance IsLeftCancelMul.contravariant_mul_le_of_contravariant_mul_lt [Mul N] [IsLeftCancelMul N] - [PartialOrder N] [ContravariantClass N N (· * ·) (· < ·)] : - ContravariantClass N N (· * ·) (· ≤ ·) where +instance IsLeftCancelMul.mulLeftReflectLE_of_mulLeftReflectLT [Mul N] [IsLeftCancelMul N] + [PartialOrder N] [MulLeftReflectLT N] : + MulLeftReflectLE N where elim := (contravariant_le_iff_contravariant_lt_and_eq N N _).mpr ⟨ContravariantClass.elim, fun _ ↦ mul_left_cancel⟩ @[to_additive] -instance IsRightCancelMul.contravariant_swap_mul_le_of_contravariant_swap_mul_lt - [Mul N] [IsRightCancelMul N] [PartialOrder N] [ContravariantClass N N (swap (· * ·)) (· < ·)] : - ContravariantClass N N (swap (· * ·)) (· ≤ ·) where +instance IsRightCancelMul.mulRightReflectLE_of_mulRightReflectLT + [Mul N] [IsRightCancelMul N] [PartialOrder N] [MulRightReflectLT N] : + MulRightReflectLE N where elim := (contravariant_le_iff_contravariant_lt_and_eq N N _).mpr ⟨ContravariantClass.elim, fun _ ↦ mul_right_cancel⟩ diff --git a/Mathlib/Algebra/Order/Monoid/Unbundled/ExistsOfLE.lean b/Mathlib/Algebra/Order/Monoid/Unbundled/ExistsOfLE.lean index bfca4ed38d9a9..4868b5b043691 100644 --- a/Mathlib/Algebra/Order/Monoid/Unbundled/ExistsOfLE.lean +++ b/Mathlib/Algebra/Order/Monoid/Unbundled/ExistsOfLE.lean @@ -44,20 +44,20 @@ instance (priority := 100) Group.existsMulOfLE (α : Type u) [Group α] [LE α] section MulOneClass variable [MulOneClass α] [Preorder α] [ExistsMulOfLE α] {a b : α} -@[to_additive] lemma exists_one_le_mul_of_le [ContravariantClass α α (· * ·) (· ≤ ·)] (h : a ≤ b) : +@[to_additive] lemma exists_one_le_mul_of_le [MulLeftReflectLE α] (h : a ≤ b) : ∃ c, 1 ≤ c ∧ a * c = b := by obtain ⟨c, rfl⟩ := exists_mul_of_le h; exact ⟨c, one_le_of_le_mul_right h, rfl⟩ -@[to_additive] lemma exists_one_lt_mul_of_lt' [ContravariantClass α α (· * ·) (· < ·)] (h : a < b) : +@[to_additive] lemma exists_one_lt_mul_of_lt' [MulLeftReflectLT α] (h : a < b) : ∃ c, 1 < c ∧ a * c = b := by obtain ⟨c, rfl⟩ := exists_mul_of_le h.le; exact ⟨c, one_lt_of_lt_mul_right h, rfl⟩ -@[to_additive] lemma le_iff_exists_one_le_mul [CovariantClass α α (· * ·) (· ≤ ·)] - [ContravariantClass α α (· * ·) (· ≤ ·)] : a ≤ b ↔ ∃ c, 1 ≤ c ∧ a * c = b := +@[to_additive] lemma le_iff_exists_one_le_mul [MulLeftMono α] + [MulLeftReflectLE α] : a ≤ b ↔ ∃ c, 1 ≤ c ∧ a * c = b := ⟨exists_one_le_mul_of_le, by rintro ⟨c, hc, rfl⟩; exact le_mul_of_one_le_right' hc⟩ -@[to_additive] lemma lt_iff_exists_one_lt_mul [CovariantClass α α (· * ·) (· < ·)] - [ContravariantClass α α (· * ·) (· < ·)] : a < b ↔ ∃ c, 1 < c ∧ a * c = b := +@[to_additive] lemma lt_iff_exists_one_lt_mul [MulLeftStrictMono α] + [MulLeftReflectLT α] : a < b ↔ ∃ c, 1 < c ∧ a * c = b := ⟨exists_one_lt_mul_of_lt', by rintro ⟨c, hc, rfl⟩; exact lt_mul_of_one_lt_right' _ hc⟩ end MulOneClass @@ -65,7 +65,7 @@ end MulOneClass section ExistsMulOfLE variable [LinearOrder α] [DenselyOrdered α] [Monoid α] [ExistsMulOfLE α] - [ContravariantClass α α (· * ·) (· < ·)] {a b : α} + [MulLeftReflectLT α] {a b : α} @[to_additive] theorem le_of_forall_one_lt_le_mul (h : ∀ ε : α, 1 < ε → a ≤ b * ε) : a ≤ b := @@ -78,7 +78,7 @@ theorem le_of_forall_one_lt_lt_mul' (h : ∀ ε : α, 1 < ε → a < b * ε) : a le_of_forall_one_lt_le_mul fun ε hε => (h ε hε).le @[to_additive] -theorem le_iff_forall_one_lt_lt_mul' [CovariantClass α α (· * ·) (· < ·)] : +theorem le_iff_forall_one_lt_lt_mul' [MulLeftStrictMono α] : a ≤ b ↔ ∀ ε, 1 < ε → a < b * ε := ⟨fun h _ => lt_mul_of_le_of_one_lt h, le_of_forall_one_lt_lt_mul'⟩ diff --git a/Mathlib/Algebra/Order/Monoid/Unbundled/MinMax.lean b/Mathlib/Algebra/Order/Monoid/Unbundled/MinMax.lean index b509125489cc7..f8cdd57c1e851 100644 --- a/Mathlib/Algebra/Order/Monoid/Unbundled/MinMax.lean +++ b/Mathlib/Algebra/Order/Monoid/Unbundled/MinMax.lean @@ -48,7 +48,7 @@ variable [Mul α] section Left -variable [CovariantClass α α (· * ·) (· ≤ ·)] +variable [MulLeftMono α] @[to_additive] theorem min_mul_mul_left (a b c : α) : min (a * b) (a * c) = a * min b c := @@ -62,7 +62,7 @@ end Left section Right -variable [CovariantClass α α (Function.swap (· * ·)) (· ≤ ·)] +variable [MulRightMono α] @[to_additive] theorem min_mul_mul_right (a b c : α) : min (a * c) (b * c) = min a b * c := @@ -75,37 +75,33 @@ theorem max_mul_mul_right (a b c : α) : max (a * c) (b * c) = max a b * c := end Right @[to_additive] -theorem lt_or_lt_of_mul_lt_mul [CovariantClass α α (· * ·) (· ≤ ·)] - [CovariantClass α α (Function.swap (· * ·)) (· ≤ ·)] {a₁ a₂ b₁ b₂ : α} : +theorem lt_or_lt_of_mul_lt_mul [MulLeftMono α] [MulRightMono α] {a₁ a₂ b₁ b₂ : α} : a₁ * b₁ < a₂ * b₂ → a₁ < a₂ ∨ b₁ < b₂ := by contrapose! exact fun h => mul_le_mul' h.1 h.2 @[to_additive] -theorem le_or_lt_of_mul_le_mul [CovariantClass α α (· * ·) (· ≤ ·)] - [CovariantClass α α (Function.swap (· * ·)) (· < ·)] {a₁ a₂ b₁ b₂ : α} : +theorem le_or_lt_of_mul_le_mul [MulLeftMono α] [MulRightStrictMono α] {a₁ a₂ b₁ b₂ : α} : a₁ * b₁ ≤ a₂ * b₂ → a₁ ≤ a₂ ∨ b₁ < b₂ := by contrapose! exact fun h => mul_lt_mul_of_lt_of_le h.1 h.2 @[to_additive] -theorem lt_or_le_of_mul_le_mul [CovariantClass α α (· * ·) (· < ·)] - [CovariantClass α α (Function.swap (· * ·)) (· ≤ ·)] {a₁ a₂ b₁ b₂ : α} : +theorem lt_or_le_of_mul_le_mul [MulLeftStrictMono α] [MulRightMono α] {a₁ a₂ b₁ b₂ : α} : a₁ * b₁ ≤ a₂ * b₂ → a₁ < a₂ ∨ b₁ ≤ b₂ := by contrapose! exact fun h => mul_lt_mul_of_le_of_lt h.1 h.2 @[to_additive] -theorem le_or_le_of_mul_le_mul [CovariantClass α α (· * ·) (· < ·)] - [CovariantClass α α (Function.swap (· * ·)) (· < ·)] {a₁ a₂ b₁ b₂ : α} : +theorem le_or_le_of_mul_le_mul [MulLeftStrictMono α] [MulRightStrictMono α] {a₁ a₂ b₁ b₂ : α} : a₁ * b₁ ≤ a₂ * b₂ → a₁ ≤ a₂ ∨ b₁ ≤ b₂ := by contrapose! exact fun h => mul_lt_mul_of_lt_of_lt h.1 h.2 @[to_additive] -theorem mul_lt_mul_iff_of_le_of_le [CovariantClass α α (· * ·) (· ≤ ·)] - [CovariantClass α α (Function.swap (· * ·)) (· ≤ ·)] [CovariantClass α α (· * ·) (· < ·)] - [CovariantClass α α (Function.swap (· * ·)) (· < ·)] {a₁ a₂ b₁ b₂ : α} (ha : a₁ ≤ a₂) +theorem mul_lt_mul_iff_of_le_of_le [MulLeftMono α] + [MulRightMono α] [MulLeftStrictMono α] + [MulRightStrictMono α] {a₁ a₂ b₁ b₂ : α} (ha : a₁ ≤ a₂) (hb : b₁ ≤ b₂) : a₁ * b₁ < a₂ * b₂ ↔ a₁ < a₂ ∨ b₁ < b₂ := by refine ⟨lt_or_lt_of_mul_lt_mul, fun h => ?_⟩ rcases h with ha' | hb' @@ -117,18 +113,17 @@ end Mul variable [MulOneClass α] @[to_additive] -theorem min_le_mul_of_one_le_right [CovariantClass α α (· * ·) (· ≤ ·)] {a b : α} (hb : 1 ≤ b) : +theorem min_le_mul_of_one_le_right [MulLeftMono α] {a b : α} (hb : 1 ≤ b) : min a b ≤ a * b := min_le_iff.2 <| Or.inl <| le_mul_of_one_le_right' hb @[to_additive] -theorem min_le_mul_of_one_le_left [CovariantClass α α (Function.swap (· * ·)) (· ≤ ·)] {a b : α} - (ha : 1 ≤ a) : min a b ≤ a * b := +theorem min_le_mul_of_one_le_left [MulRightMono α] {a b : α} (ha : 1 ≤ a) : + min a b ≤ a * b := min_le_iff.2 <| Or.inr <| le_mul_of_one_le_left' ha @[to_additive] -theorem max_le_mul_of_one_le [CovariantClass α α (· * ·) (· ≤ ·)] - [CovariantClass α α (Function.swap (· * ·)) (· ≤ ·)] {a b : α} (ha : 1 ≤ a) (hb : 1 ≤ b) : +theorem max_le_mul_of_one_le [MulLeftMono α] [MulRightMono α] {a b : α} (ha : 1 ≤ a) (hb : 1 ≤ b) : max a b ≤ a * b := max_le_iff.2 ⟨le_mul_of_one_le_right' hb, le_mul_of_one_le_left' ha⟩ diff --git a/Mathlib/Algebra/Order/Monoid/Unbundled/OrderDual.lean b/Mathlib/Algebra/Order/Monoid/Unbundled/OrderDual.lean index 7a9819c7230f1..c93cbe11c4902 100644 --- a/Mathlib/Algebra/Order/Monoid/Unbundled/OrderDual.lean +++ b/Mathlib/Algebra/Order/Monoid/Unbundled/OrderDual.lean @@ -17,47 +17,35 @@ open Function namespace OrderDual @[to_additive] -instance contravariantClass_mul_le [LE α] [Mul α] [c : ContravariantClass α α (· * ·) (· ≤ ·)] : - ContravariantClass αᵒᵈ αᵒᵈ (· * ·) (· ≤ ·) := +instance mulLeftReflectLE [LE α] [Mul α] [c : MulLeftReflectLE α] : MulLeftReflectLE αᵒᵈ := ⟨c.1.flip⟩ @[to_additive] -instance covariantClass_mul_le [LE α] [Mul α] [c : CovariantClass α α (· * ·) (· ≤ ·)] : - CovariantClass αᵒᵈ αᵒᵈ (· * ·) (· ≤ ·) := +instance mulLeftMono [LE α] [Mul α] [c : MulLeftMono α] : MulLeftMono αᵒᵈ := ⟨c.1.flip⟩ @[to_additive] -instance contravariantClass_swap_mul_le [LE α] [Mul α] - [c : ContravariantClass α α (swap (· * ·)) (· ≤ ·)] : - ContravariantClass αᵒᵈ αᵒᵈ (swap (· * ·)) (· ≤ ·) := +instance mulRightReflectLE [LE α] [Mul α] [c : MulRightReflectLE α] : MulRightReflectLE αᵒᵈ := ⟨c.1.flip⟩ @[to_additive] -instance covariantClass_swap_mul_le [LE α] [Mul α] - [c : CovariantClass α α (swap (· * ·)) (· ≤ ·)] : - CovariantClass αᵒᵈ αᵒᵈ (swap (· * ·)) (· ≤ ·) := +instance mulRightMono [LE α] [Mul α] [c : MulRightMono α] : MulRightMono αᵒᵈ := ⟨c.1.flip⟩ @[to_additive] -instance contravariantClass_mul_lt [LT α] [Mul α] [c : ContravariantClass α α (· * ·) (· < ·)] : - ContravariantClass αᵒᵈ αᵒᵈ (· * ·) (· < ·) := +instance mulLeftReflectLT [LT α] [Mul α] [c : MulLeftReflectLT α] : MulLeftReflectLT αᵒᵈ := ⟨c.1.flip⟩ @[to_additive] -instance covariantClass_mul_lt [LT α] [Mul α] [c : CovariantClass α α (· * ·) (· < ·)] : - CovariantClass αᵒᵈ αᵒᵈ (· * ·) (· < ·) := +instance mulLeftStrictMono [LT α] [Mul α] [c : MulLeftStrictMono α] : MulLeftStrictMono αᵒᵈ := ⟨c.1.flip⟩ @[to_additive] -instance contravariantClass_swap_mul_lt [LT α] [Mul α] - [c : ContravariantClass α α (swap (· * ·)) (· < ·)] : - ContravariantClass αᵒᵈ αᵒᵈ (swap (· * ·)) (· < ·) := +instance mulRightReflectLT [LT α] [Mul α] [c : MulRightReflectLT α] : MulRightReflectLT αᵒᵈ := ⟨c.1.flip⟩ @[to_additive] -instance covariantClass_swap_mul_lt [LT α] [Mul α] - [c : CovariantClass α α (swap (· * ·)) (· < ·)] : - CovariantClass αᵒᵈ αᵒᵈ (swap (· * ·)) (· < ·) := +instance mulRightStrictMono [LT α] [Mul α] [c : MulRightStrictMono α] : MulRightStrictMono αᵒᵈ := ⟨c.1.flip⟩ end OrderDual diff --git a/Mathlib/Algebra/Order/Monoid/Unbundled/Pow.lean b/Mathlib/Algebra/Order/Monoid/Unbundled/Pow.lean index eb892cfeda6b8..348b449849781 100644 --- a/Mathlib/Algebra/Order/Monoid/Unbundled/Pow.lean +++ b/Mathlib/Algebra/Order/Monoid/Unbundled/Pow.lean @@ -26,7 +26,7 @@ variable [Preorder M] namespace Left -variable [CovariantClass M M (· * ·) (· ≤ ·)] {a : M} +variable [MulLeftMono M] {a : M} @[to_additive Left.nsmul_nonneg] theorem one_le_pow_of_le (ha : 1 ≤ a) : ∀ n : ℕ, 1 ≤ a ^ n @@ -58,7 +58,7 @@ end Left section Left -variable [CovariantClass M M (· * ·) (· ≤ ·)] +variable [MulLeftMono M] @[to_additive nsmul_left_monotone] theorem pow_right_monotone {a : M} (ha : 1 ≤ a) : Monotone fun n : ℕ ↦ a ^ n := @@ -80,7 +80,7 @@ end Left section LeftLt -variable [CovariantClass M M (· * ·) (· < ·)] {a : M} {n m : ℕ} +variable [MulLeftStrictMono M] {a : M} {n m : ℕ} @[to_additive nsmul_left_strictMono] theorem pow_right_strictMono' (ha : 1 < a) : StrictMono ((a ^ ·) : ℕ → M) := @@ -94,7 +94,7 @@ end LeftLt section Right -variable [CovariantClass M M (swap (· * ·)) (· ≤ ·)] {x : M} +variable [MulRightMono M] {x : M} @[to_additive Right.nsmul_nonneg] theorem Right.one_le_pow_of_le (hx : 1 ≤ x) : ∀ {n : ℕ}, 1 ≤ x ^ n @@ -123,8 +123,7 @@ end Right section CovariantLTSwap -variable [Preorder β] [CovariantClass M M (· * ·) (· < ·)] - [CovariantClass M M (swap (· * ·)) (· < ·)] {f : β → M} {n : ℕ} +variable [Preorder β] [MulLeftStrictMono M] [MulRightStrictMono M] {f : β → M} {n : ℕ} @[to_additive StrictMono.const_nsmul] theorem StrictMono.pow_const (hf : StrictMono f) : ∀ {n : ℕ}, n ≠ 0 → StrictMono (f · ^ n) @@ -145,8 +144,7 @@ end CovariantLTSwap section CovariantLESwap -variable [Preorder β] [CovariantClass M M (· * ·) (· ≤ ·)] - [CovariantClass M M (swap (· * ·)) (· ≤ ·)] +variable [Preorder β] [MulLeftMono M] [MulRightMono M] @[to_additive (attr := mono, gcongr) nsmul_le_nsmul_right] theorem pow_le_pow_left' {a b : M} (hab : a ≤ b) : ∀ i : ℕ, a ^ i ≤ b ^ i @@ -175,7 +173,7 @@ variable [LinearOrder M] section CovariantLE -variable [CovariantClass M M (· * ·) (· ≤ ·)] +variable [MulLeftMono M] -- This generalises to lattices. See `pow_two_semiclosed` @[to_additive nsmul_nonneg_iff] @@ -203,7 +201,7 @@ end CovariantLE section CovariantLT -variable [CovariantClass M M (· * ·) (· < ·)] {a : M} {m n : ℕ} +variable [MulLeftStrictMono M] {a : M} {m n : ℕ} @[to_additive nsmul_le_nsmul_iff_left] theorem pow_le_pow_iff_right' (ha : 1 < a) : a ^ m ≤ a ^ n ↔ m ≤ n := @@ -217,7 +215,7 @@ end CovariantLT section CovariantLESwap -variable [CovariantClass M M (· * ·) (· ≤ ·)] [CovariantClass M M (swap (· * ·)) (· ≤ ·)] +variable [MulLeftMono M] [MulRightMono M] @[to_additive lt_of_nsmul_lt_nsmul_right] theorem lt_of_pow_lt_pow_left' {a b : M} (n : ℕ) : a ^ n < b ^ n → a < b := @@ -235,7 +233,7 @@ end CovariantLESwap section CovariantLTSwap -variable [CovariantClass M M (· * ·) (· < ·)] [CovariantClass M M (swap (· * ·)) (· < ·)] +variable [MulLeftStrictMono M] [MulRightStrictMono M] @[to_additive le_of_nsmul_le_nsmul_right] theorem le_of_pow_le_pow_left' {a b : M} {n : ℕ} (hn : n ≠ 0) : a ^ n ≤ b ^ n → a ≤ b := @@ -252,18 +250,18 @@ theorem le_max_of_sq_le_mul {a b c : M} (h : a ^ 2 ≤ b * c) : a ≤ max b c := end CovariantLTSwap @[to_additive Left.nsmul_neg_iff] -theorem Left.pow_lt_one_iff' [CovariantClass M M (· * ·) (· < ·)] {n : ℕ} {x : M} (hn : 0 < n) : +theorem Left.pow_lt_one_iff' [MulLeftStrictMono M] {n : ℕ} {x : M} (hn : 0 < n) : x ^ n < 1 ↔ x < 1 := - haveI := covariantClass_le_of_lt M M (· * ·) + haveI := mulLeftMono_of_mulLeftStrictMono M pow_lt_one_iff hn.ne' -theorem Left.pow_lt_one_iff [CovariantClass M M (· * ·) (· < ·)] {n : ℕ} {x : M} (hn : 0 < n) : +theorem Left.pow_lt_one_iff [MulLeftStrictMono M] {n : ℕ} {x : M} (hn : 0 < n) : x ^ n < 1 ↔ x < 1 := Left.pow_lt_one_iff' hn @[to_additive] -theorem Right.pow_lt_one_iff [CovariantClass M M (swap (· * ·)) (· < ·)] {n : ℕ} {x : M} +theorem Right.pow_lt_one_iff [MulRightStrictMono M] {n : ℕ} {x : M} (hn : 0 < n) : x ^ n < 1 ↔ x < 1 := - haveI := covariantClass_le_of_lt M M (swap (· * ·)) + haveI := mulRightMono_of_mulRightStrictMono M ⟨fun H => not_le.mp fun k => H.not_le <| Right.one_le_pow_of_le k, Right.pow_lt_one_of_lt hn⟩ end LinearOrder @@ -272,7 +270,7 @@ end Monoid section DivInvMonoid -variable [DivInvMonoid G] [Preorder G] [CovariantClass G G (· * ·) (· ≤ ·)] +variable [DivInvMonoid G] [Preorder G] [MulLeftMono G] @[to_additive zsmul_nonneg] theorem one_le_zpow {x : G} (H : 1 ≤ x) {n : ℤ} (hn : 0 ≤ n) : 1 ≤ x ^ n := by diff --git a/Mathlib/Algebra/Order/Monoid/Unbundled/WithTop.lean b/Mathlib/Algebra/Order/Monoid/Unbundled/WithTop.lean index b7cbbd3985728..f23cb20dc1b90 100644 --- a/Mathlib/Algebra/Order/Monoid/Unbundled/WithTop.lean +++ b/Mathlib/Algebra/Order/Monoid/Unbundled/WithTop.lean @@ -143,22 +143,19 @@ theorem add_left_cancel_iff [IsLeftCancelAdd α] (ha : a ≠ ⊤) : a + b = a + theorem add_left_cancel [IsLeftCancelAdd α] (ha : a ≠ ⊤) (h : a + b = a + c) : b = c := (WithTop.add_left_cancel_iff ha).1 h -instance covariantClass_add_le [LE α] [CovariantClass α α (· + ·) (· ≤ ·)] : - CovariantClass (WithTop α) (WithTop α) (· + ·) (· ≤ ·) := +instance addLeftMono [LE α] [AddLeftMono α] : AddLeftMono (WithTop α) := ⟨fun a b c h => by cases a <;> cases c <;> try exact le_top rcases le_coe_iff.1 h with ⟨b, rfl, _⟩ exact coe_le_coe.2 (add_le_add_left (coe_le_coe.1 h) _)⟩ -instance covariantClass_swap_add_le [LE α] [CovariantClass α α (swap (· + ·)) (· ≤ ·)] : - CovariantClass (WithTop α) (WithTop α) (swap (· + ·)) (· ≤ ·) := +instance addRightMono [LE α] [AddRightMono α] : AddRightMono (WithTop α) := ⟨fun a b c h => by cases a <;> cases c <;> try exact le_top rcases le_coe_iff.1 h with ⟨b, rfl, _⟩ exact coe_le_coe.2 (add_le_add_right (coe_le_coe.1 h) _)⟩ -instance contravariantClass_add_lt [LT α] [ContravariantClass α α (· + ·) (· < ·)] : - ContravariantClass (WithTop α) (WithTop α) (· + ·) (· < ·) := +instance addLeftReflectLT [LT α] [AddLeftReflectLT α] : AddLeftReflectLT (WithTop α) := ⟨fun a b c h => by induction a; · exact (WithTop.not_top_lt _ h).elim induction b; · exact (WithTop.not_top_lt _ h).elim @@ -166,15 +163,14 @@ instance contravariantClass_add_lt [LT α] [ContravariantClass α α (· + ·) ( · exact coe_lt_top _ · exact coe_lt_coe.2 (lt_of_add_lt_add_left <| coe_lt_coe.1 h)⟩ -instance contravariantClass_swap_add_lt [LT α] [ContravariantClass α α (swap (· + ·)) (· < ·)] : - ContravariantClass (WithTop α) (WithTop α) (swap (· + ·)) (· < ·) := +instance addRightReflectLT [LT α] [AddRightReflectLT α] : AddRightReflectLT (WithTop α) := ⟨fun a b c h => by cases a <;> cases b <;> try exact (WithTop.not_top_lt _ h).elim cases c · exact coe_lt_top _ · exact coe_lt_coe.2 (lt_of_add_lt_add_right <| coe_lt_coe.1 h)⟩ -protected theorem le_of_add_le_add_left [LE α] [ContravariantClass α α (· + ·) (· ≤ ·)] (ha : a ≠ ⊤) +protected theorem le_of_add_le_add_left [LE α] [AddLeftReflectLE α] (ha : a ≠ ⊤) (h : a + b ≤ a + c) : b ≤ c := by lift a to α using ha induction c @@ -184,7 +180,7 @@ protected theorem le_of_add_le_add_left [LE α] [ContravariantClass α α (· + · simp only [← coe_add, coe_le_coe] at h ⊢ exact le_of_add_le_add_left h -protected theorem le_of_add_le_add_right [LE α] [ContravariantClass α α (swap (· + ·)) (· ≤ ·)] +protected theorem le_of_add_le_add_right [LE α] [AddRightReflectLE α] (ha : a ≠ ⊤) (h : b + a ≤ c + a) : b ≤ c := by lift a to α using ha cases c @@ -193,7 +189,7 @@ protected theorem le_of_add_le_add_right [LE α] [ContravariantClass α α (swap · exact (not_top_le_coe _ h).elim · exact coe_le_coe.2 (le_of_add_le_add_right <| coe_le_coe.1 h) -protected theorem add_lt_add_left [LT α] [CovariantClass α α (· + ·) (· < ·)] (ha : a ≠ ⊤) +protected theorem add_lt_add_left [LT α] [AddLeftStrictMono α] (ha : a ≠ ⊤) (h : b < c) : a + b < a + c := by lift a to α using ha rcases lt_iff_exists_coe.1 h with ⟨b, rfl, h'⟩ @@ -201,7 +197,7 @@ protected theorem add_lt_add_left [LT α] [CovariantClass α α (· + ·) (· < · exact coe_lt_top _ · exact coe_lt_coe.2 (add_lt_add_left (coe_lt_coe.1 h) _) -protected theorem add_lt_add_right [LT α] [CovariantClass α α (swap (· + ·)) (· < ·)] (ha : a ≠ ⊤) +protected theorem add_lt_add_right [LT α] [AddRightStrictMono α] (ha : a ≠ ⊤) (h : b < c) : b + a < c + a := by lift a to α using ha rcases lt_iff_exists_coe.1 h with ⟨b, rfl, h'⟩ @@ -209,29 +205,29 @@ protected theorem add_lt_add_right [LT α] [CovariantClass α α (swap (· + ·) · exact coe_lt_top _ · exact coe_lt_coe.2 (add_lt_add_right (coe_lt_coe.1 h) _) -protected theorem add_le_add_iff_left [LE α] [CovariantClass α α (· + ·) (· ≤ ·)] - [ContravariantClass α α (· + ·) (· ≤ ·)] (ha : a ≠ ⊤) : a + b ≤ a + c ↔ b ≤ c := +protected theorem add_le_add_iff_left [LE α] [AddLeftMono α] + [AddLeftReflectLE α] (ha : a ≠ ⊤) : a + b ≤ a + c ↔ b ≤ c := ⟨WithTop.le_of_add_le_add_left ha, fun h => add_le_add_left h a⟩ -protected theorem add_le_add_iff_right [LE α] [CovariantClass α α (swap (· + ·)) (· ≤ ·)] - [ContravariantClass α α (swap (· + ·)) (· ≤ ·)] (ha : a ≠ ⊤) : b + a ≤ c + a ↔ b ≤ c := +protected theorem add_le_add_iff_right [LE α] [AddRightMono α] + [AddRightReflectLE α] (ha : a ≠ ⊤) : b + a ≤ c + a ↔ b ≤ c := ⟨WithTop.le_of_add_le_add_right ha, fun h => add_le_add_right h a⟩ -protected theorem add_lt_add_iff_left [LT α] [CovariantClass α α (· + ·) (· < ·)] - [ContravariantClass α α (· + ·) (· < ·)] (ha : a ≠ ⊤) : a + b < a + c ↔ b < c := +protected theorem add_lt_add_iff_left [LT α] [AddLeftStrictMono α] + [AddLeftReflectLT α] (ha : a ≠ ⊤) : a + b < a + c ↔ b < c := ⟨lt_of_add_lt_add_left, WithTop.add_lt_add_left ha⟩ -protected theorem add_lt_add_iff_right [LT α] [CovariantClass α α (swap (· + ·)) (· < ·)] - [ContravariantClass α α (swap (· + ·)) (· < ·)] (ha : a ≠ ⊤) : b + a < c + a ↔ b < c := +protected theorem add_lt_add_iff_right [LT α] [AddRightStrictMono α] + [AddRightReflectLT α] (ha : a ≠ ⊤) : b + a < c + a ↔ b < c := ⟨lt_of_add_lt_add_right, WithTop.add_lt_add_right ha⟩ -protected theorem add_lt_add_of_le_of_lt [Preorder α] [CovariantClass α α (· + ·) (· < ·)] - [CovariantClass α α (swap (· + ·)) (· ≤ ·)] (ha : a ≠ ⊤) (hab : a ≤ b) (hcd : c < d) : +protected theorem add_lt_add_of_le_of_lt [Preorder α] [AddLeftStrictMono α] + [AddRightMono α] (ha : a ≠ ⊤) (hab : a ≤ b) (hcd : c < d) : a + c < b + d := (WithTop.add_lt_add_left ha hcd).trans_le <| add_le_add_right hab _ -protected theorem add_lt_add_of_lt_of_le [Preorder α] [CovariantClass α α (· + ·) (· ≤ ·)] - [CovariantClass α α (swap (· + ·)) (· < ·)] (hc : c ≠ ⊤) (hab : a < b) (hcd : c ≤ d) : +protected theorem add_lt_add_of_lt_of_le [Preorder α] [AddLeftMono α] + [AddRightStrictMono α] (hc : c ≠ ⊤) (hab : a < b) (hcd : c ≤ d) : a + c < b + d := (WithTop.add_lt_add_right hc hab).trans_le <| add_le_add_left hcd _ @@ -598,61 +594,57 @@ protected def _root_.AddMonoidHom.withBotMap {M N : Type*} [AddZeroClass M] [Add variable [Preorder α] -instance covariantClass_add_le [CovariantClass α α (· + ·) (· ≤ ·)] : - CovariantClass (WithBot α) (WithBot α) (· + ·) (· ≤ ·) := - OrderDual.covariantClass_add_le (α := WithTop αᵒᵈ) +instance addLeftMono [AddLeftMono α] : AddLeftMono (WithBot α) := + OrderDual.addLeftMono (α := WithTop αᵒᵈ) -instance covariantClass_swap_add_le [CovariantClass α α (swap (· + ·)) (· ≤ ·)] : - CovariantClass (WithBot α) (WithBot α) (swap (· + ·)) (· ≤ ·) := - OrderDual.covariantClass_swap_add_le (α := WithTop αᵒᵈ) +instance addRightMono [AddRightMono α] : AddRightMono (WithBot α) := + OrderDual.addRightMono (α := WithTop αᵒᵈ) -instance contravariantClass_add_lt [ContravariantClass α α (· + ·) (· < ·)] : - ContravariantClass (WithBot α) (WithBot α) (· + ·) (· < ·) := - OrderDual.contravariantClass_add_lt (α := WithTop αᵒᵈ) +instance addLeftReflectLT [AddLeftReflectLT α] : AddLeftReflectLT (WithBot α) := + OrderDual.addLeftReflectLT (α := WithTop αᵒᵈ) -instance contravariantClass_swap_add_lt [ContravariantClass α α (swap (· + ·)) (· < ·)] : - ContravariantClass (WithBot α) (WithBot α) (swap (· + ·)) (· < ·) := - OrderDual.contravariantClass_swap_add_lt (α := WithTop αᵒᵈ) +instance addRightReflectLT [AddRightReflectLT α] : AddRightReflectLT (WithBot α) := + OrderDual.addRightReflectLT (α := WithTop αᵒᵈ) -protected theorem le_of_add_le_add_left [ContravariantClass α α (· + ·) (· ≤ ·)] (ha : a ≠ ⊥) +protected theorem le_of_add_le_add_left [AddLeftReflectLE α] (ha : a ≠ ⊥) (h : a + b ≤ a + c) : b ≤ c := WithTop.le_of_add_le_add_left (α := αᵒᵈ) ha h -protected theorem le_of_add_le_add_right [ContravariantClass α α (swap (· + ·)) (· ≤ ·)] +protected theorem le_of_add_le_add_right [AddRightReflectLE α] (ha : a ≠ ⊥) (h : b + a ≤ c + a) : b ≤ c := WithTop.le_of_add_le_add_right (α := αᵒᵈ) ha h -protected theorem add_lt_add_left [CovariantClass α α (· + ·) (· < ·)] (ha : a ≠ ⊥) (h : b < c) : +protected theorem add_lt_add_left [AddLeftStrictMono α] (ha : a ≠ ⊥) (h : b < c) : a + b < a + c := WithTop.add_lt_add_left (α := αᵒᵈ) ha h -protected theorem add_lt_add_right [CovariantClass α α (swap (· + ·)) (· < ·)] (ha : a ≠ ⊥) +protected theorem add_lt_add_right [AddRightStrictMono α] (ha : a ≠ ⊥) (h : b < c) : b + a < c + a := WithTop.add_lt_add_right (α := αᵒᵈ) ha h -protected theorem add_le_add_iff_left [CovariantClass α α (· + ·) (· ≤ ·)] - [ContravariantClass α α (· + ·) (· ≤ ·)] (ha : a ≠ ⊥) : a + b ≤ a + c ↔ b ≤ c := +protected theorem add_le_add_iff_left [AddLeftMono α] + [AddLeftReflectLE α] (ha : a ≠ ⊥) : a + b ≤ a + c ↔ b ≤ c := ⟨WithBot.le_of_add_le_add_left ha, fun h => add_le_add_left h a⟩ -protected theorem add_le_add_iff_right [CovariantClass α α (swap (· + ·)) (· ≤ ·)] - [ContravariantClass α α (swap (· + ·)) (· ≤ ·)] (ha : a ≠ ⊥) : b + a ≤ c + a ↔ b ≤ c := +protected theorem add_le_add_iff_right [AddRightMono α] + [AddRightReflectLE α] (ha : a ≠ ⊥) : b + a ≤ c + a ↔ b ≤ c := ⟨WithBot.le_of_add_le_add_right ha, fun h => add_le_add_right h a⟩ -protected theorem add_lt_add_iff_left [CovariantClass α α (· + ·) (· < ·)] - [ContravariantClass α α (· + ·) (· < ·)] (ha : a ≠ ⊥) : a + b < a + c ↔ b < c := +protected theorem add_lt_add_iff_left [AddLeftStrictMono α] + [AddLeftReflectLT α] (ha : a ≠ ⊥) : a + b < a + c ↔ b < c := ⟨lt_of_add_lt_add_left, WithBot.add_lt_add_left ha⟩ -protected theorem add_lt_add_iff_right [CovariantClass α α (swap (· + ·)) (· < ·)] - [ContravariantClass α α (swap (· + ·)) (· < ·)] (ha : a ≠ ⊥) : b + a < c + a ↔ b < c := +protected theorem add_lt_add_iff_right [AddRightStrictMono α] + [AddRightReflectLT α] (ha : a ≠ ⊥) : b + a < c + a ↔ b < c := ⟨lt_of_add_lt_add_right, WithBot.add_lt_add_right ha⟩ -protected theorem add_lt_add_of_le_of_lt [CovariantClass α α (· + ·) (· < ·)] - [CovariantClass α α (swap (· + ·)) (· ≤ ·)] (hb : b ≠ ⊥) (hab : a ≤ b) (hcd : c < d) : +protected theorem add_lt_add_of_le_of_lt [AddLeftStrictMono α] + [AddRightMono α] (hb : b ≠ ⊥) (hab : a ≤ b) (hcd : c < d) : a + c < b + d := WithTop.add_lt_add_of_le_of_lt (α := αᵒᵈ) hb hab hcd -protected theorem add_lt_add_of_lt_of_le [CovariantClass α α (· + ·) (· ≤ ·)] - [CovariantClass α α (swap (· + ·)) (· < ·)] (hd : d ≠ ⊥) (hab : a < b) (hcd : c ≤ d) : +protected theorem add_lt_add_of_lt_of_le [AddLeftMono α] + [AddRightStrictMono α] (hd : d ≠ ⊥) (hab : a < b) (hcd : c ≤ d) : a + c < b + d := WithTop.add_lt_add_of_lt_of_le (α := αᵒᵈ) hd hab hcd diff --git a/Mathlib/Algebra/Order/Positive/Ring.lean b/Mathlib/Algebra/Order/Positive/Ring.lean index 98d4be8046d53..ff395b0df05ef 100644 --- a/Mathlib/Algebra/Order/Positive/Ring.lean +++ b/Mathlib/Algebra/Order/Positive/Ring.lean @@ -23,7 +23,7 @@ variable {M R : Type*} section AddBasic -variable [AddMonoid M] [Preorder M] [CovariantClass M M (· + ·) (· < ·)] +variable [AddMonoid M] [Preorder M] [AddLeftStrictMono M] instance : Add { x : M // 0 < x } := ⟨fun x y => ⟨x + y, add_pos x.2 y.2⟩⟩ @@ -36,46 +36,39 @@ instance addSemigroup : AddSemigroup { x : M // 0 < x } := Subtype.coe_injective.addSemigroup _ coe_add instance addCommSemigroup {M : Type*} [AddCommMonoid M] [Preorder M] - [CovariantClass M M (· + ·) (· < ·)] : AddCommSemigroup { x : M // 0 < x } := + [AddLeftStrictMono M] : AddCommSemigroup { x : M // 0 < x } := Subtype.coe_injective.addCommSemigroup _ coe_add instance addLeftCancelSemigroup {M : Type*} [AddLeftCancelMonoid M] [Preorder M] - [CovariantClass M M (· + ·) (· < ·)] : AddLeftCancelSemigroup { x : M // 0 < x } := + [AddLeftStrictMono M] : AddLeftCancelSemigroup { x : M // 0 < x } := Subtype.coe_injective.addLeftCancelSemigroup _ coe_add instance addRightCancelSemigroup {M : Type*} [AddRightCancelMonoid M] [Preorder M] - [CovariantClass M M (· + ·) (· < ·)] : AddRightCancelSemigroup { x : M // 0 < x } := + [AddLeftStrictMono M] : AddRightCancelSemigroup { x : M // 0 < x } := Subtype.coe_injective.addRightCancelSemigroup _ coe_add -instance covariantClass_add_lt : - CovariantClass { x : M // 0 < x } { x : M // 0 < x } (· + ·) (· < ·) := +instance addLeftStrictMono : AddLeftStrictMono { x : M // 0 < x } := ⟨fun _ y z hyz => Subtype.coe_lt_coe.1 <| add_lt_add_left (show (y : M) < z from hyz) _⟩ -instance covariantClass_swap_add_lt [CovariantClass M M (swap (· + ·)) (· < ·)] : - CovariantClass { x : M // 0 < x } { x : M // 0 < x } (swap (· + ·)) (· < ·) := +instance addRightStrictMono [AddRightStrictMono M] : AddRightStrictMono { x : M // 0 < x } := ⟨fun _ y z hyz => Subtype.coe_lt_coe.1 <| add_lt_add_right (show (y : M) < z from hyz) _⟩ -instance contravariantClass_add_lt [ContravariantClass M M (· + ·) (· < ·)] : - ContravariantClass { x : M // 0 < x } { x : M // 0 < x } (· + ·) (· < ·) := +instance addLeftReflectLT [AddLeftReflectLT M] : AddLeftReflectLT { x : M // 0 < x } := ⟨fun _ _ _ h => Subtype.coe_lt_coe.1 <| lt_of_add_lt_add_left h⟩ -instance contravariantClass_swap_add_lt [ContravariantClass M M (swap (· + ·)) (· < ·)] : - ContravariantClass { x : M // 0 < x } { x : M // 0 < x } (swap (· + ·)) (· < ·) := +instance addRightReflectLT [AddRightReflectLT M] : AddRightReflectLT { x : M // 0 < x } := ⟨fun _ _ _ h => Subtype.coe_lt_coe.1 <| lt_of_add_lt_add_right h⟩ -instance contravariantClass_add_le [ContravariantClass M M (· + ·) (· ≤ ·)] : - ContravariantClass { x : M // 0 < x } { x : M // 0 < x } (· + ·) (· ≤ ·) := +instance addLeftReflectLE [AddLeftReflectLE M] : AddLeftReflectLE { x : M // 0 < x } := ⟨fun _ _ _ h => Subtype.coe_le_coe.1 <| le_of_add_le_add_left h⟩ -instance contravariantClass_swap_add_le [ContravariantClass M M (swap (· + ·)) (· ≤ ·)] : - ContravariantClass { x : M // 0 < x } { x : M // 0 < x } (swap (· + ·)) (· ≤ ·) := +instance addRightReflectLE [AddRightReflectLE M] : AddRightReflectLE { x : M // 0 < x } := ⟨fun _ _ _ h => Subtype.coe_le_coe.1 <| le_of_add_le_add_right h⟩ end AddBasic -instance covariantClass_add_le [AddMonoid M] [PartialOrder M] - [CovariantClass M M (· + ·) (· < ·)] : - CovariantClass { x : M // 0 < x } { x : M // 0 < x } (· + ·) (· ≤ ·) := +instance addLeftMono [AddMonoid M] [PartialOrder M] [AddLeftStrictMono M] : + AddLeftMono { x : M // 0 < x } := ⟨@fun _ _ _ h₁ => StrictMono.monotone (fun _ _ h => add_lt_add_left h _) h₁⟩ section Mul diff --git a/Mathlib/Algebra/Order/Ring/Basic.lean b/Mathlib/Algebra/Order/Ring/Basic.lean index 0dac73109c9d0..284772133e61f 100644 --- a/Mathlib/Algebra/Order/Ring/Basic.lean +++ b/Mathlib/Algebra/Order/Ring/Basic.lean @@ -21,7 +21,7 @@ variable {α M R : Type*} namespace MonoidHom -variable [Ring R] [Monoid M] [LinearOrder M] [CovariantClass M M (· * ·) (· ≤ ·)] (f : R →* M) +variable [Ring R] [Monoid M] [LinearOrder M] [MulLeftMono M] (f : R →* M) theorem map_neg_one : f (-1) = 1 := (pow_eq_one_iff (Nat.succ_ne_zero 1)).1 <| by rw [← map_pow, neg_one_sq, map_one] diff --git a/Mathlib/Algebra/Order/Ring/Canonical.lean b/Mathlib/Algebra/Order/Ring/Canonical.lean index 3c450c7ad0c1a..3c7a907838251 100644 --- a/Mathlib/Algebra/Order/Ring/Canonical.lean +++ b/Mathlib/Algebra/Order/Ring/Canonical.lean @@ -51,8 +51,8 @@ instance (priority := 100) toNoZeroDivisors : NoZeroDivisors α := ⟨CanonicallyOrderedCommSemiring.eq_zero_or_eq_zero_of_mul_eq_zero⟩ -- see Note [lower instance priority] -instance (priority := 100) toCovariantClassMulLE : CovariantClass α α (· * ·) (· ≤ ·) := by - refine ⟨fun a b c h => ?_⟩ +instance (priority := 100) toMulLeftMono : MulLeftMono α := by + refine ⟨fun a b c h => ?_⟩; dsimp rcases exists_add_of_le h with ⟨c, rfl⟩ rw [mul_add] apply self_le_add_right @@ -106,7 +106,7 @@ protected theorem tsub_mul (h : AddLECancellable (b * c)) : (a - b) * c = a * c end AddLECancellable -variable [ContravariantClass α α (· + ·) (· ≤ ·)] +variable [AddLeftReflectLE α] theorem mul_tsub (a b c : α) : a * (b - c) = a * b - a * c := Contravariant.AddLECancellable.mul_tsub diff --git a/Mathlib/Algebra/Order/Ring/Cast.lean b/Mathlib/Algebra/Order/Ring/Cast.lean index bd47e1f894415..74581d5714f35 100644 --- a/Mathlib/Algebra/Order/Ring/Cast.lean +++ b/Mathlib/Algebra/Order/Ring/Cast.lean @@ -27,7 +27,7 @@ variable {R : Type*} namespace Int section OrderedAddCommGroupWithOne -variable [AddCommGroupWithOne R] [PartialOrder R] [CovariantClass R R (· + ·) (· ≤ ·)] +variable [AddCommGroupWithOne R] [PartialOrder R] [AddLeftMono R] variable [ZeroLEOneClass R] lemma cast_mono : Monotone (Int.cast : ℤ → R) := by diff --git a/Mathlib/Algebra/Order/Ring/Unbundled/Basic.lean b/Mathlib/Algebra/Order/Ring/Unbundled/Basic.lean index 9c0bbff7ea773..06cc0315f08b0 100644 --- a/Mathlib/Algebra/Order/Ring/Unbundled/Basic.lean +++ b/Mathlib/Algebra/Order/Ring/Unbundled/Basic.lean @@ -126,7 +126,7 @@ variable {α : Type u} {β : Type*} `zero_le_one` field. -/ -theorem add_one_le_two_mul [LE α] [Semiring α] [CovariantClass α α (· + ·) (· ≤ ·)] {a : α} +theorem add_one_le_two_mul [LE α] [Semiring α] [AddLeftMono α] {a : α} (a1 : 1 ≤ a) : a + 1 ≤ 2 * a := calc a + 1 ≤ a + a := add_le_add_left a1 a @@ -137,7 +137,7 @@ section OrderedSemiring variable [Semiring α] [Preorder α] {a b c d : α} -- Porting note: it's unfortunate we need to write `(@one_le_two α)` here. -theorem add_le_mul_two_add [ZeroLEOneClass α] [MulPosMono α] [CovariantClass α α (· + ·) (· ≤ ·)] +theorem add_le_mul_two_add [ZeroLEOneClass α] [MulPosMono α] [AddLeftMono α] (a2 : 2 ≤ a) (b0 : 0 ≤ b) : a + (2 + b) ≤ a * (2 + b) := calc a + (2 + b) ≤ a + (a + a * b) := @@ -145,7 +145,7 @@ theorem add_le_mul_two_add [ZeroLEOneClass α] [MulPosMono α] [CovariantClass _ ≤ a * (2 + b) := by rw [mul_add, mul_two, add_assoc] theorem mul_le_mul_of_nonpos_left [ExistsAddOfLE α] [PosMulMono α] - [CovariantClass α α (swap (· + ·)) (· ≤ ·)] [ContravariantClass α α (swap (· + ·)) (· ≤ ·)] + [AddRightMono α] [AddRightReflectLE α] (h : b ≤ a) (hc : c ≤ 0) : c * a ≤ c * b := by obtain ⟨d, hcd⟩ := exists_add_of_le hc refine le_of_add_le_add_right (a := d * b + d * a) ?_ @@ -155,7 +155,7 @@ theorem mul_le_mul_of_nonpos_left [ExistsAddOfLE α] [PosMulMono α] _ = _ := by rw [← add_assoc, ← add_mul, ← hcd, zero_mul, zero_add] theorem mul_le_mul_of_nonpos_right [ExistsAddOfLE α] [MulPosMono α] - [CovariantClass α α (swap (· + ·)) (· ≤ ·)] [ContravariantClass α α (swap (· + ·)) (· ≤ ·)] + [AddRightMono α] [AddRightReflectLE α] (h : b ≤ a) (hc : c ≤ 0) : a * c ≤ b * c := by obtain ⟨d, hcd⟩ := exists_add_of_le hc refine le_of_add_le_add_right (a := b * d + a * d) ?_ @@ -165,61 +165,61 @@ theorem mul_le_mul_of_nonpos_right [ExistsAddOfLE α] [MulPosMono α] _ = _ := by rw [← add_assoc, ← mul_add, ← hcd, mul_zero, zero_add] theorem mul_nonneg_of_nonpos_of_nonpos [ExistsAddOfLE α] [MulPosMono α] - [CovariantClass α α (swap (· + ·)) (· ≤ ·)] [ContravariantClass α α (swap (· + ·)) (· ≤ ·)] + [AddRightMono α] [AddRightReflectLE α] (ha : a ≤ 0) (hb : b ≤ 0) : 0 ≤ a * b := by simpa only [zero_mul] using mul_le_mul_of_nonpos_right ha hb theorem mul_le_mul_of_nonneg_of_nonpos [ExistsAddOfLE α] [MulPosMono α] [PosMulMono α] - [CovariantClass α α (swap (· + ·)) (· ≤ ·)] [ContravariantClass α α (swap (· + ·)) (· ≤ ·)] + [AddRightMono α] [AddRightReflectLE α] (hca : c ≤ a) (hbd : b ≤ d) (hc : 0 ≤ c) (hb : b ≤ 0) : a * b ≤ c * d := (mul_le_mul_of_nonpos_right hca hb).trans <| mul_le_mul_of_nonneg_left hbd hc theorem mul_le_mul_of_nonneg_of_nonpos' [ExistsAddOfLE α] [PosMulMono α] [MulPosMono α] - [CovariantClass α α (swap (· + ·)) (· ≤ ·)] [ContravariantClass α α (swap (· + ·)) (· ≤ ·)] + [AddRightMono α] [AddRightReflectLE α] (hca : c ≤ a) (hbd : b ≤ d) (ha : 0 ≤ a) (hd : d ≤ 0) : a * b ≤ c * d := (mul_le_mul_of_nonneg_left hbd ha).trans <| mul_le_mul_of_nonpos_right hca hd theorem mul_le_mul_of_nonpos_of_nonneg [ExistsAddOfLE α] [MulPosMono α] [PosMulMono α] - [CovariantClass α α (swap (· + ·)) (· ≤ ·)] [ContravariantClass α α (swap (· + ·)) (· ≤ ·)] + [AddRightMono α] [AddRightReflectLE α] (hac : a ≤ c) (hdb : d ≤ b) (hc : c ≤ 0) (hb : 0 ≤ b) : a * b ≤ c * d := (mul_le_mul_of_nonneg_right hac hb).trans <| mul_le_mul_of_nonpos_left hdb hc theorem mul_le_mul_of_nonpos_of_nonneg' [ExistsAddOfLE α] [PosMulMono α] [MulPosMono α] - [CovariantClass α α (swap (· + ·)) (· ≤ ·)] [ContravariantClass α α (swap (· + ·)) (· ≤ ·)] + [AddRightMono α] [AddRightReflectLE α] (hca : c ≤ a) (hbd : b ≤ d) (ha : 0 ≤ a) (hd : d ≤ 0) : a * b ≤ c * d := (mul_le_mul_of_nonneg_left hbd ha).trans <| mul_le_mul_of_nonpos_right hca hd theorem mul_le_mul_of_nonpos_of_nonpos [ExistsAddOfLE α] [MulPosMono α] [PosMulMono α] - [CovariantClass α α (swap (· + ·)) (· ≤ ·)] [ContravariantClass α α (swap (· + ·)) (· ≤ ·)] + [AddRightMono α] [AddRightReflectLE α] (hca : c ≤ a) (hdb : d ≤ b) (hc : c ≤ 0) (hb : b ≤ 0) : a * b ≤ c * d := (mul_le_mul_of_nonpos_right hca hb).trans <| mul_le_mul_of_nonpos_left hdb hc theorem mul_le_mul_of_nonpos_of_nonpos' [ExistsAddOfLE α] [PosMulMono α] [MulPosMono α] - [CovariantClass α α (swap (· + ·)) (· ≤ ·)] [ContravariantClass α α (swap (· + ·)) (· ≤ ·)] + [AddRightMono α] [AddRightReflectLE α] (hca : c ≤ a) (hdb : d ≤ b) (ha : a ≤ 0) (hd : d ≤ 0) : a * b ≤ c * d := (mul_le_mul_of_nonpos_left hdb ha).trans <| mul_le_mul_of_nonpos_right hca hd /-- Variant of `mul_le_of_le_one_left` for `b` non-positive instead of non-negative. -/ theorem le_mul_of_le_one_left [ExistsAddOfLE α] [MulPosMono α] - [CovariantClass α α (swap (· + ·)) (· ≤ ·)] [ContravariantClass α α (swap (· + ·)) (· ≤ ·)] + [AddRightMono α] [AddRightReflectLE α] (hb : b ≤ 0) (h : a ≤ 1) : b ≤ a * b := by simpa only [one_mul] using mul_le_mul_of_nonpos_right h hb /-- Variant of `le_mul_of_one_le_left` for `b` non-positive instead of non-negative. -/ theorem mul_le_of_one_le_left [ExistsAddOfLE α] [MulPosMono α] - [CovariantClass α α (swap (· + ·)) (· ≤ ·)] [ContravariantClass α α (swap (· + ·)) (· ≤ ·)] + [AddRightMono α] [AddRightReflectLE α] (hb : b ≤ 0) (h : 1 ≤ a) : a * b ≤ b := by simpa only [one_mul] using mul_le_mul_of_nonpos_right h hb /-- Variant of `mul_le_of_le_one_right` for `a` non-positive instead of non-negative. -/ theorem le_mul_of_le_one_right [ExistsAddOfLE α] [PosMulMono α] - [CovariantClass α α (swap (· + ·)) (· ≤ ·)] [ContravariantClass α α (swap (· + ·)) (· ≤ ·)] + [AddRightMono α] [AddRightReflectLE α] (ha : a ≤ 0) (h : b ≤ 1) : a ≤ a * b := by simpa only [mul_one] using mul_le_mul_of_nonpos_left h ha /-- Variant of `le_mul_of_one_le_right` for `a` non-positive instead of non-negative. -/ theorem mul_le_of_one_le_right [ExistsAddOfLE α] [PosMulMono α] - [CovariantClass α α (swap (· + ·)) (· ≤ ·)] [ContravariantClass α α (swap (· + ·)) (· ≤ ·)] + [AddRightMono α] [AddRightReflectLE α] (ha : a ≤ 0) (h : 1 ≤ b) : a * b ≤ a := by simpa only [mul_one] using mul_le_mul_of_nonpos_left h ha @@ -228,49 +228,49 @@ section Monotone variable [Preorder β] {f g : β → α} theorem antitone_mul_left [ExistsAddOfLE α] [PosMulMono α] - [CovariantClass α α (swap (· + ·)) (· ≤ ·)] [ContravariantClass α α (swap (· + ·)) (· ≤ ·)] + [AddRightMono α] [AddRightReflectLE α] {a : α} (ha : a ≤ 0) : Antitone (a * ·) := fun _ _ b_le_c => mul_le_mul_of_nonpos_left b_le_c ha theorem antitone_mul_right [ExistsAddOfLE α] [MulPosMono α] - [CovariantClass α α (swap (· + ·)) (· ≤ ·)] [ContravariantClass α α (swap (· + ·)) (· ≤ ·)] + [AddRightMono α] [AddRightReflectLE α] {a : α} (ha : a ≤ 0) : Antitone fun x => x * a := fun _ _ b_le_c => mul_le_mul_of_nonpos_right b_le_c ha theorem Monotone.const_mul_of_nonpos [ExistsAddOfLE α] [PosMulMono α] - [CovariantClass α α (swap (· + ·)) (· ≤ ·)] [ContravariantClass α α (swap (· + ·)) (· ≤ ·)] + [AddRightMono α] [AddRightReflectLE α] (hf : Monotone f) (ha : a ≤ 0) : Antitone fun x => a * f x := (antitone_mul_left ha).comp_monotone hf theorem Monotone.mul_const_of_nonpos [ExistsAddOfLE α] [MulPosMono α] - [CovariantClass α α (swap (· + ·)) (· ≤ ·)] [ContravariantClass α α (swap (· + ·)) (· ≤ ·)] + [AddRightMono α] [AddRightReflectLE α] (hf : Monotone f) (ha : a ≤ 0) : Antitone fun x => f x * a := (antitone_mul_right ha).comp_monotone hf theorem Antitone.const_mul_of_nonpos [ExistsAddOfLE α] [PosMulMono α] - [CovariantClass α α (swap (· + ·)) (· ≤ ·)] [ContravariantClass α α (swap (· + ·)) (· ≤ ·)] + [AddRightMono α] [AddRightReflectLE α] (hf : Antitone f) (ha : a ≤ 0) : Monotone fun x => a * f x := (antitone_mul_left ha).comp hf theorem Antitone.mul_const_of_nonpos [ExistsAddOfLE α] [MulPosMono α] - [CovariantClass α α (swap (· + ·)) (· ≤ ·)] [ContravariantClass α α (swap (· + ·)) (· ≤ ·)] + [AddRightMono α] [AddRightReflectLE α] (hf : Antitone f) (ha : a ≤ 0) : Monotone fun x => f x * a := (antitone_mul_right ha).comp hf theorem Antitone.mul_monotone [ExistsAddOfLE α] [PosMulMono α] [MulPosMono α] - [CovariantClass α α (swap (· + ·)) (· ≤ ·)] [ContravariantClass α α (swap (· + ·)) (· ≤ ·)] + [AddRightMono α] [AddRightReflectLE α] (hf : Antitone f) (hg : Monotone g) (hf₀ : ∀ x, f x ≤ 0) (hg₀ : ∀ x, 0 ≤ g x) : Antitone (f * g) := fun _ _ h => mul_le_mul_of_nonpos_of_nonneg (hf h) (hg h) (hf₀ _) (hg₀ _) theorem Monotone.mul_antitone [ExistsAddOfLE α] [PosMulMono α] [MulPosMono α] - [CovariantClass α α (swap (· + ·)) (· ≤ ·)] [ContravariantClass α α (swap (· + ·)) (· ≤ ·)] + [AddRightMono α] [AddRightReflectLE α] (hf : Monotone f) (hg : Antitone g) (hf₀ : ∀ x, 0 ≤ f x) (hg₀ : ∀ x, g x ≤ 0) : Antitone (f * g) := fun _ _ h => mul_le_mul_of_nonneg_of_nonpos (hf h) (hg h) (hf₀ _) (hg₀ _) theorem Antitone.mul [ExistsAddOfLE α] [PosMulMono α] [MulPosMono α] - [CovariantClass α α (swap (· + ·)) (· ≤ ·)] [ContravariantClass α α (swap (· + ·)) (· ≤ ·)] + [AddRightMono α] [AddRightReflectLE α] (hf : Antitone f) (hg : Antitone g) (hf₀ : ∀ x, f x ≤ 0) (hg₀ : ∀ x, g x ≤ 0) : Monotone (f * g) := fun _ _ h => mul_le_mul_of_nonpos_of_nonpos (hf h) (hg h) (hf₀ _) (hg₀ _) @@ -284,11 +284,11 @@ section StrictOrderedSemiring variable [Semiring α] [PartialOrder α] {a b c d : α} theorem lt_two_mul_self [ZeroLEOneClass α] [MulPosStrictMono α] [NeZero (R := α) 1] - [CovariantClass α α (· + ·) (· < ·)] (ha : 0 < a) : a < 2 * a := + [AddLeftStrictMono α] (ha : 0 < a) : a < 2 * a := lt_mul_of_one_lt_left ha one_lt_two theorem mul_lt_mul_of_neg_left [ExistsAddOfLE α] [PosMulStrictMono α] - [CovariantClass α α (swap (· + ·)) (· < ·)] [ContravariantClass α α (swap (· + ·)) (· < ·)] + [AddRightStrictMono α] [AddRightReflectLT α] (h : b < a) (hc : c < 0) : c * a < c * b := by obtain ⟨d, hcd⟩ := exists_add_of_le hc.le refine (add_lt_add_iff_right (d * b + d * a)).1 ?_ @@ -298,7 +298,7 @@ theorem mul_lt_mul_of_neg_left [ExistsAddOfLE α] [PosMulStrictMono α] _ = _ := by rw [← add_assoc, ← add_mul, ← hcd, zero_mul, zero_add] theorem mul_lt_mul_of_neg_right [ExistsAddOfLE α] [MulPosStrictMono α] - [CovariantClass α α (swap (· + ·)) (· < ·)] [ContravariantClass α α (swap (· + ·)) (· < ·)] + [AddRightStrictMono α] [AddRightReflectLT α] (h : b < a) (hc : c < 0) : a * c < b * c := by obtain ⟨d, hcd⟩ := exists_add_of_le hc.le refine (add_lt_add_iff_right (b * d + a * d)).1 ?_ @@ -308,31 +308,31 @@ theorem mul_lt_mul_of_neg_right [ExistsAddOfLE α] [MulPosStrictMono α] _ = _ := by rw [← add_assoc, ← mul_add, ← hcd, mul_zero, zero_add] theorem mul_pos_of_neg_of_neg [ExistsAddOfLE α] [MulPosStrictMono α] - [CovariantClass α α (swap (· + ·)) (· < ·)] [ContravariantClass α α (swap (· + ·)) (· < ·)] + [AddRightStrictMono α] [AddRightReflectLT α] {a b : α} (ha : a < 0) (hb : b < 0) : 0 < a * b := by simpa only [zero_mul] using mul_lt_mul_of_neg_right ha hb /-- Variant of `mul_lt_of_lt_one_left` for `b` negative instead of positive. -/ theorem lt_mul_of_lt_one_left [ExistsAddOfLE α] [MulPosStrictMono α] - [CovariantClass α α (swap (· + ·)) (· < ·)] [ContravariantClass α α (swap (· + ·)) (· < ·)] + [AddRightStrictMono α] [AddRightReflectLT α] (hb : b < 0) (h : a < 1) : b < a * b := by simpa only [one_mul] using mul_lt_mul_of_neg_right h hb /-- Variant of `lt_mul_of_one_lt_left` for `b` negative instead of positive. -/ theorem mul_lt_of_one_lt_left [ExistsAddOfLE α] [MulPosStrictMono α] - [CovariantClass α α (swap (· + ·)) (· < ·)] [ContravariantClass α α (swap (· + ·)) (· < ·)] + [AddRightStrictMono α] [AddRightReflectLT α] (hb : b < 0) (h : 1 < a) : a * b < b := by simpa only [one_mul] using mul_lt_mul_of_neg_right h hb /-- Variant of `mul_lt_of_lt_one_right` for `a` negative instead of positive. -/ theorem lt_mul_of_lt_one_right [ExistsAddOfLE α] [PosMulStrictMono α] - [CovariantClass α α (swap (· + ·)) (· < ·)] [ContravariantClass α α (swap (· + ·)) (· < ·)] + [AddRightStrictMono α] [AddRightReflectLT α] (ha : a < 0) (h : b < 1) : a < a * b := by simpa only [mul_one] using mul_lt_mul_of_neg_left h ha /-- Variant of `lt_mul_of_lt_one_right` for `a` negative instead of positive. -/ theorem mul_lt_of_one_lt_right [ExistsAddOfLE α] [PosMulStrictMono α] - [CovariantClass α α (swap (· + ·)) (· < ·)] [ContravariantClass α α (swap (· + ·)) (· < ·)] + [AddRightStrictMono α] [AddRightReflectLT α] (ha : a < 0) (h : 1 < b) : a * b < a := by simpa only [mul_one] using mul_lt_mul_of_neg_left h ha @@ -341,32 +341,32 @@ section Monotone variable [Preorder β] {f : β → α} theorem strictAnti_mul_left [ExistsAddOfLE α] [PosMulStrictMono α] - [CovariantClass α α (swap (· + ·)) (· < ·)] [ContravariantClass α α (swap (· + ·)) (· < ·)] + [AddRightStrictMono α] [AddRightReflectLT α] {a : α} (ha : a < 0) : StrictAnti (a * ·) := fun _ _ b_lt_c => mul_lt_mul_of_neg_left b_lt_c ha theorem strictAnti_mul_right [ExistsAddOfLE α] [MulPosStrictMono α] - [CovariantClass α α (swap (· + ·)) (· < ·)] [ContravariantClass α α (swap (· + ·)) (· < ·)] + [AddRightStrictMono α] [AddRightReflectLT α] {a : α} (ha : a < 0) : StrictAnti fun x => x * a := fun _ _ b_lt_c => mul_lt_mul_of_neg_right b_lt_c ha theorem StrictMono.const_mul_of_neg [ExistsAddOfLE α] [PosMulStrictMono α] - [CovariantClass α α (swap (· + ·)) (· < ·)] [ContravariantClass α α (swap (· + ·)) (· < ·)] + [AddRightStrictMono α] [AddRightReflectLT α] (hf : StrictMono f) (ha : a < 0) : StrictAnti fun x => a * f x := (strictAnti_mul_left ha).comp_strictMono hf theorem StrictMono.mul_const_of_neg [ExistsAddOfLE α] [MulPosStrictMono α] - [CovariantClass α α (swap (· + ·)) (· < ·)] [ContravariantClass α α (swap (· + ·)) (· < ·)] + [AddRightStrictMono α] [AddRightReflectLT α] (hf : StrictMono f) (ha : a < 0) : StrictAnti fun x => f x * a := (strictAnti_mul_right ha).comp_strictMono hf theorem StrictAnti.const_mul_of_neg [ExistsAddOfLE α] [PosMulStrictMono α] - [CovariantClass α α (swap (· + ·)) (· < ·)] [ContravariantClass α α (swap (· + ·)) (· < ·)] + [AddRightStrictMono α] [AddRightReflectLT α] (hf : StrictAnti f) (ha : a < 0) : StrictMono fun x => a * f x := (strictAnti_mul_left ha).comp hf theorem StrictAnti.mul_const_of_neg [ExistsAddOfLE α] [MulPosStrictMono α] - [CovariantClass α α (swap (· + ·)) (· < ·)] [ContravariantClass α α (swap (· + ·)) (· < ·)] + [AddRightStrictMono α] [AddRightReflectLT α] (hf : StrictAnti f) (ha : a < 0) : StrictMono fun x => f x * a := (strictAnti_mul_right ha).comp hf @@ -374,7 +374,7 @@ end Monotone /-- Binary **rearrangement inequality**. -/ lemma mul_add_mul_le_mul_add_mul [ExistsAddOfLE α] [MulPosMono α] - [CovariantClass α α (· + ·) (· ≤ ·)] [ContravariantClass α α (· + ·) (· ≤ ·)] + [AddLeftMono α] [AddLeftReflectLE α] (hab : a ≤ b) (hcd : c ≤ d) : a * d + b * c ≤ a * c + b * d := by obtain ⟨b, rfl⟩ := exists_add_of_le hab obtain ⟨d, hd, rfl⟩ := exists_nonneg_add_of_le hcd @@ -383,15 +383,15 @@ lemma mul_add_mul_le_mul_add_mul [ExistsAddOfLE α] [MulPosMono α] /-- Binary **rearrangement inequality**. -/ lemma mul_add_mul_le_mul_add_mul' [ExistsAddOfLE α] [MulPosMono α] - [CovariantClass α α (· + ·) (· ≤ ·)] [ContravariantClass α α (· + ·) (· ≤ ·)] + [AddLeftMono α] [AddLeftReflectLE α] (hba : b ≤ a) (hdc : d ≤ c) : a * d + b * c ≤ a * c + b * d := by rw [add_comm (a * d), add_comm (a * c)]; exact mul_add_mul_le_mul_add_mul hba hdc -variable [ContravariantClass α α (· + ·) (· < ·)] +variable [AddLeftReflectLT α] /-- Binary strict **rearrangement inequality**. -/ lemma mul_add_mul_lt_mul_add_mul [ExistsAddOfLE α] [MulPosStrictMono α] - [CovariantClass α α (· + ·) (· < ·)] + [AddLeftStrictMono α] (hab : a < b) (hcd : c < d) : a * d + b * c < a * c + b * d := by obtain ⟨b, rfl⟩ := exists_add_of_le hab.le obtain ⟨d, hd, rfl⟩ := exists_pos_add_of_lt' hcd @@ -400,7 +400,7 @@ lemma mul_add_mul_lt_mul_add_mul [ExistsAddOfLE α] [MulPosStrictMono α] /-- Binary **rearrangement inequality**. -/ lemma mul_add_mul_lt_mul_add_mul' [ExistsAddOfLE α] [MulPosStrictMono α] - [CovariantClass α α (· + ·) (· < ·)] + [AddLeftStrictMono α] (hba : b < a) (hdc : d < c) : a * d + b * c < a * c + b * d := by rw [add_comm (a * d), add_comm (a * c)] exact mul_add_mul_lt_mul_add_mul hba hdc @@ -449,7 +449,7 @@ theorem mul_nonneg_iff_of_pos_right [MulPosStrictMono α] simpa using (mul_le_mul_right h : 0 * c ≤ b * c ↔ 0 ≤ b) theorem add_le_mul_of_left_le_right [ZeroLEOneClass α] [NeZero (R := α) 1] - [MulPosStrictMono α] [CovariantClass α α (· + ·) (· ≤ ·)] + [MulPosStrictMono α] [AddLeftMono α] (a2 : 2 ≤ a) (ab : a ≤ b) : a + b ≤ a * b := have : 0 < b := calc @@ -463,7 +463,7 @@ theorem add_le_mul_of_left_le_right [ZeroLEOneClass α] [NeZero (R := α) 1] -- Porting note: we used to not need the type annotation on `(0 : α)` at the start of the `calc`. theorem add_le_mul_of_right_le_left [ZeroLEOneClass α] [NeZero (R := α) 1] - [CovariantClass α α (· + ·) (· ≤ ·)] [PosMulStrictMono α] + [AddLeftMono α] [PosMulStrictMono α] (b2 : 2 ≤ b) (ba : b ≤ a) : a + b ≤ a * b := have : 0 < a := calc (0 : α) @@ -476,13 +476,13 @@ theorem add_le_mul_of_right_le_left [ZeroLEOneClass α] [NeZero (R := α) 1] _ ≤ a * b := (mul_le_mul_left this).mpr b2 theorem add_le_mul [ZeroLEOneClass α] [NeZero (R := α) 1] - [MulPosStrictMono α] [PosMulStrictMono α] [CovariantClass α α (· + ·) (· ≤ ·)] + [MulPosStrictMono α] [PosMulStrictMono α] [AddLeftMono α] (a2 : 2 ≤ a) (b2 : 2 ≤ b) : a + b ≤ a * b := if hab : a ≤ b then add_le_mul_of_left_le_right a2 hab else add_le_mul_of_right_le_left b2 (le_of_not_le hab) theorem add_le_mul' [ZeroLEOneClass α] [NeZero (R := α) 1] - [MulPosStrictMono α] [PosMulStrictMono α] [CovariantClass α α (· + ·) (· ≤ ·)] + [MulPosStrictMono α] [PosMulStrictMono α] [AddLeftMono α] (a2 : 2 ≤ a) (b2 : 2 ≤ b) : a + b ≤ b * a := (le_of_eq (add_comm _ _)).trans (add_le_mul b2 a2) @@ -570,20 +570,20 @@ lemma sign_cases_of_C_mul_pow_nonneg [PosMulStrictMono α] simpa only [pow_one] using h 1 theorem mul_pos_iff [ExistsAddOfLE α] [PosMulStrictMono α] [MulPosStrictMono α] - [CovariantClass α α (· + ·) (· < ·)] [ContravariantClass α α (· + ·) (· < ·)] : + [AddLeftStrictMono α] [AddLeftReflectLT α] : 0 < a * b ↔ 0 < a ∧ 0 < b ∨ a < 0 ∧ b < 0 := ⟨pos_and_pos_or_neg_and_neg_of_mul_pos, fun h => h.elim (and_imp.2 mul_pos) (and_imp.2 mul_pos_of_neg_of_neg)⟩ theorem mul_nonneg_iff [ExistsAddOfLE α] [MulPosStrictMono α] [PosMulStrictMono α] - [ContravariantClass α α (· + ·) (· ≤ ·)] [CovariantClass α α (· + ·) (· ≤ ·)]: + [AddLeftReflectLE α] [AddLeftMono α]: 0 ≤ a * b ↔ 0 ≤ a ∧ 0 ≤ b ∨ a ≤ 0 ∧ b ≤ 0 := ⟨nonneg_and_nonneg_or_nonpos_and_nonpos_of_mul_nonneg, fun h => h.elim (and_imp.2 mul_nonneg) (and_imp.2 mul_nonneg_of_nonpos_of_nonpos)⟩ /-- Out of three elements of a `LinearOrderedRing`, two must have the same sign. -/ theorem mul_nonneg_of_three [ExistsAddOfLE α] [MulPosStrictMono α] [PosMulStrictMono α] - [CovariantClass α α (· + ·) (· ≤ ·)] [ContravariantClass α α (· + ·) (· ≤ ·)] + [AddLeftMono α] [AddLeftReflectLE α] (a b c : α) : 0 ≤ a * b ∨ 0 ≤ b * c ∨ 0 ≤ c * a := by iterate 3 rw [mul_nonneg_iff] have or_a := le_total 0 a @@ -609,7 +609,7 @@ theorem mul_nonneg_of_three [ExistsAddOfLE α] [MulPosStrictMono α] [PosMulStri (fun (h6 : a ≤ 0) => Or.inl (Or.inr ⟨h6, h4⟩)))) lemma mul_nonneg_iff_pos_imp_nonneg [ExistsAddOfLE α] [PosMulStrictMono α] [MulPosStrictMono α] - [CovariantClass α α (· + ·) (· ≤ ·)] [ContravariantClass α α (· + ·) (· ≤ ·)] : + [AddLeftMono α] [AddLeftReflectLE α] : 0 ≤ a * b ↔ (0 < a → 0 ≤ b) ∧ (0 < b → 0 ≤ a) := by refine mul_nonneg_iff.trans ?_ simp_rw [← not_le, ← or_iff_not_imp_left] @@ -619,51 +619,51 @@ lemma mul_nonneg_iff_pos_imp_nonneg [ExistsAddOfLE α] [PosMulStrictMono α] [Mu @[simp] theorem mul_le_mul_left_of_neg [ExistsAddOfLE α] [PosMulStrictMono α] - [CovariantClass α α (swap (· + ·)) (· ≤ ·)] [ContravariantClass α α (swap (· + ·)) (· ≤ ·)] + [AddRightMono α] [AddRightReflectLE α] {a b c : α} (h : c < 0) : c * a ≤ c * b ↔ b ≤ a := (strictAnti_mul_left h).le_iff_le @[simp] theorem mul_le_mul_right_of_neg [ExistsAddOfLE α] [MulPosStrictMono α] - [CovariantClass α α (swap (· + ·)) (· ≤ ·)] [ContravariantClass α α (swap (· + ·)) (· ≤ ·)] + [AddRightMono α] [AddRightReflectLE α] {a b c : α} (h : c < 0) : a * c ≤ b * c ↔ b ≤ a := (strictAnti_mul_right h).le_iff_le @[simp] theorem mul_lt_mul_left_of_neg [ExistsAddOfLE α] [PosMulStrictMono α] - [CovariantClass α α (swap (· + ·)) (· < ·)] [ContravariantClass α α (swap (· + ·)) (· < ·)] + [AddRightStrictMono α] [AddRightReflectLT α] {a b c : α} (h : c < 0) : c * a < c * b ↔ b < a := (strictAnti_mul_left h).lt_iff_lt @[simp] theorem mul_lt_mul_right_of_neg [ExistsAddOfLE α] [MulPosStrictMono α] - [CovariantClass α α (swap (· + ·)) (· < ·)] [ContravariantClass α α (swap (· + ·)) (· < ·)] + [AddRightStrictMono α] [AddRightReflectLT α] {a b c : α} (h : c < 0) : a * c < b * c ↔ b < a := (strictAnti_mul_right h).lt_iff_lt theorem lt_of_mul_lt_mul_of_nonpos_left [ExistsAddOfLE α] [PosMulMono α] - [CovariantClass α α (swap (· + ·)) (· ≤ ·)] [ContravariantClass α α (swap (· + ·)) (· ≤ ·)] + [AddRightMono α] [AddRightReflectLE α] (h : c * a < c * b) (hc : c ≤ 0) : b < a := (antitone_mul_left hc).reflect_lt h theorem lt_of_mul_lt_mul_of_nonpos_right [ExistsAddOfLE α] [MulPosMono α] - [CovariantClass α α (swap (· + ·)) (· ≤ ·)] [ContravariantClass α α (swap (· + ·)) (· ≤ ·)] + [AddRightMono α] [AddRightReflectLE α] (h : a * c < b * c) (hc : c ≤ 0) : b < a := (antitone_mul_right hc).reflect_lt h theorem cmp_mul_neg_left [ExistsAddOfLE α] [PosMulStrictMono α] - [ContravariantClass α α (swap (· + ·)) (· < ·)] [CovariantClass α α (swap (· + ·)) (· < ·)] + [AddRightReflectLT α] [AddRightStrictMono α] {a : α} (ha : a < 0) (b c : α) : cmp (a * b) (a * c) = cmp c b := (strictAnti_mul_left ha).cmp_map_eq b c theorem cmp_mul_neg_right [ExistsAddOfLE α] [MulPosStrictMono α] - [ContravariantClass α α (swap (· + ·)) (· < ·)] [CovariantClass α α (swap (· + ·)) (· < ·)] + [AddRightReflectLT α] [AddRightStrictMono α] {a : α} (ha : a < 0) (b c : α) : cmp (b * a) (c * a) = cmp c b := (strictAnti_mul_right ha).cmp_map_eq b c @[simp] theorem mul_self_pos [ExistsAddOfLE α] [PosMulStrictMono α] [MulPosStrictMono α] - [CovariantClass α α (· + ·) (· < ·)] [ContravariantClass α α (· + ·) (· < ·)] + [AddLeftStrictMono α] [AddLeftReflectLT α] {a : α} : 0 < a * a ↔ a ≠ 0 := by constructor · rintro h rfl @@ -674,37 +674,37 @@ theorem mul_self_pos [ExistsAddOfLE α] [PosMulStrictMono α] [MulPosStrictMono exacts [mul_pos_of_neg_of_neg h h, mul_pos h h] theorem nonneg_of_mul_nonpos_left [ExistsAddOfLE α] [MulPosStrictMono α] - [CovariantClass α α (swap (· + ·)) (· ≤ ·)] [ContravariantClass α α (swap (· + ·)) (· ≤ ·)] + [AddRightMono α] [AddRightReflectLE α] {a b : α} (h : a * b ≤ 0) (hb : b < 0) : 0 ≤ a := le_of_not_gt fun ha => absurd h (mul_pos_of_neg_of_neg ha hb).not_le theorem nonneg_of_mul_nonpos_right [ExistsAddOfLE α] [MulPosStrictMono α] - [CovariantClass α α (swap (· + ·)) (· ≤ ·)] [ContravariantClass α α (swap (· + ·)) (· ≤ ·)] + [AddRightMono α] [AddRightReflectLE α] {a b : α} (h : a * b ≤ 0) (ha : a < 0) : 0 ≤ b := le_of_not_gt fun hb => absurd h (mul_pos_of_neg_of_neg ha hb).not_le theorem pos_of_mul_neg_left [ExistsAddOfLE α] [MulPosMono α] - [CovariantClass α α (swap (· + ·)) (· ≤ ·)] [ContravariantClass α α (swap (· + ·)) (· ≤ ·)] + [AddRightMono α] [AddRightReflectLE α] {a b : α} (h : a * b < 0) (hb : b ≤ 0) : 0 < a := lt_of_not_ge fun ha => absurd h (mul_nonneg_of_nonpos_of_nonpos ha hb).not_lt theorem pos_of_mul_neg_right [ExistsAddOfLE α] [MulPosMono α] - [CovariantClass α α (swap (· + ·)) (· ≤ ·)] [ContravariantClass α α (swap (· + ·)) (· ≤ ·)] + [AddRightMono α] [AddRightReflectLE α] {a b : α} (h : a * b < 0) (ha : a ≤ 0) : 0 < b := lt_of_not_ge fun hb => absurd h (mul_nonneg_of_nonpos_of_nonpos ha hb).not_lt theorem neg_iff_pos_of_mul_neg [ExistsAddOfLE α] [PosMulMono α] [MulPosMono α] - [CovariantClass α α (swap (· + ·)) (· ≤ ·)] [ContravariantClass α α (swap (· + ·)) (· ≤ ·)] + [AddRightMono α] [AddRightReflectLE α] (hab : a * b < 0) : a < 0 ↔ 0 < b := ⟨pos_of_mul_neg_right hab ∘ le_of_lt, neg_of_mul_neg_left hab ∘ le_of_lt⟩ theorem pos_iff_neg_of_mul_neg [ExistsAddOfLE α] [PosMulMono α] [MulPosMono α] - [CovariantClass α α (swap (· + ·)) (· ≤ ·)] [ContravariantClass α α (swap (· + ·)) (· ≤ ·)] + [AddRightMono α] [AddRightReflectLE α] (hab : a * b < 0) : 0 < a ↔ b < 0 := ⟨neg_of_mul_neg_right hab ∘ le_of_lt, pos_of_mul_neg_left hab ∘ le_of_lt⟩ lemma sq_nonneg [IsRightCancelAdd α] - [ZeroLEOneClass α] [ExistsAddOfLE α] [PosMulMono α] [CovariantClass α α (· + ·) (· < ·)] + [ZeroLEOneClass α] [ExistsAddOfLE α] [PosMulMono α] [AddLeftStrictMono α] (a : α) : 0 ≤ a ^ 2 := by obtain ha | ha := le_total 0 a · exact pow_nonneg ha _ @@ -720,20 +720,20 @@ lemma sq_nonneg [IsRightCancelAdd α] alias pow_two_nonneg := sq_nonneg lemma mul_self_nonneg [IsRightCancelAdd α] - [ZeroLEOneClass α] [ExistsAddOfLE α] [PosMulMono α] [CovariantClass α α (· + ·) (· < ·)] + [ZeroLEOneClass α] [ExistsAddOfLE α] [PosMulMono α] [AddLeftStrictMono α] (a : α) : 0 ≤ a * a := by simpa only [sq] using sq_nonneg a /-- The sum of two squares is zero iff both elements are zero. -/ lemma mul_self_add_mul_self_eq_zero [IsRightCancelAdd α] [NoZeroDivisors α] [ZeroLEOneClass α] [ExistsAddOfLE α] [PosMulMono α] - [CovariantClass α α (· + ·) (· ≤ ·)] [CovariantClass α α (· + ·) (· < ·)] : + [AddLeftMono α] [AddLeftStrictMono α] : a * a + b * b = 0 ↔ a = 0 ∧ b = 0 := by rw [add_eq_zero_iff_of_nonneg, mul_self_eq_zero (M₀ := α), mul_self_eq_zero (M₀ := α)] <;> apply mul_self_nonneg lemma eq_zero_of_mul_self_add_mul_self_eq_zero [IsRightCancelAdd α] [NoZeroDivisors α] [ZeroLEOneClass α] [ExistsAddOfLE α] [PosMulMono α] - [CovariantClass α α (· + ·) (· ≤ ·)] [CovariantClass α α (· + ·) (· < ·)] + [AddLeftMono α] [AddLeftStrictMono α] (h : a * a + b * b = 0) : a = 0 := (mul_self_add_mul_self_eq_zero.mp h).left @@ -754,7 +754,7 @@ lemma max_mul_mul_le_max_mul_max [PosMulMono α] [MulPosMono α] (b c : α) (ha /-- Binary, squared, and division-free **arithmetic mean-geometric mean inequality** (aka AM-GM inequality) for linearly ordered commutative semirings. -/ lemma two_mul_le_add_sq [ExistsAddOfLE α] [MulPosStrictMono α] - [ContravariantClass α α (· + ·) (· ≤ ·)] [CovariantClass α α (· + ·) (· ≤ ·)] + [AddLeftReflectLE α] [AddLeftMono α] (a b : α) : 2 * a * b ≤ a ^ 2 + b ^ 2 := by simpa [fn_min_add_fn_max (fun x ↦ x * x), sq, two_mul, add_mul] using mul_add_mul_le_mul_add_mul (@min_le_max _ _ a b) (@min_le_max _ _ a b) @@ -784,42 +784,42 @@ variable [Ring α] [LinearOrder α] {a b : α} -- `[Semiring α] [LinearOrder α] [ExistsAddOfLE α] ..`? lemma mul_neg_iff [PosMulStrictMono α] [MulPosStrictMono α] - [ContravariantClass α α (· + ·) (· < ·)] [CovariantClass α α (· + ·) (· < ·)] : + [AddLeftReflectLT α] [AddLeftStrictMono α] : a * b < 0 ↔ 0 < a ∧ b < 0 ∨ a < 0 ∧ 0 < b := by rw [← neg_pos, neg_mul_eq_mul_neg, mul_pos_iff (α := α), neg_pos, neg_lt_zero] lemma mul_nonpos_iff [MulPosStrictMono α] [PosMulStrictMono α] - [ContravariantClass α α (· + ·) (· ≤ ·)] [CovariantClass α α (· + ·) (· ≤ ·)] : + [AddLeftReflectLE α] [AddLeftMono α] : a * b ≤ 0 ↔ 0 ≤ a ∧ b ≤ 0 ∨ a ≤ 0 ∧ 0 ≤ b := by rw [← neg_nonneg, neg_mul_eq_mul_neg, mul_nonneg_iff (α := α), neg_nonneg, neg_nonpos] lemma mul_nonneg_iff_neg_imp_nonpos [PosMulStrictMono α] [MulPosStrictMono α] - [CovariantClass α α (· + ·) (· ≤ ·)] [ContravariantClass α α (· + ·) (· ≤ ·)] : + [AddLeftMono α] [AddLeftReflectLE α] : 0 ≤ a * b ↔ (a < 0 → b ≤ 0) ∧ (b < 0 → a ≤ 0) := by rw [← neg_mul_neg, mul_nonneg_iff_pos_imp_nonneg (α := α)]; simp only [neg_pos, neg_nonneg] lemma mul_nonpos_iff_pos_imp_nonpos [PosMulStrictMono α] [MulPosStrictMono α] - [CovariantClass α α (· + ·) (· ≤ ·)] [ContravariantClass α α (· + ·) (· ≤ ·)] : + [AddLeftMono α] [AddLeftReflectLE α] : a * b ≤ 0 ↔ (0 < a → b ≤ 0) ∧ (b < 0 → 0 ≤ a) := by rw [← neg_nonneg, ← mul_neg, mul_nonneg_iff_pos_imp_nonneg (α := α)] simp only [neg_pos, neg_nonneg] lemma mul_nonpos_iff_neg_imp_nonneg [PosMulStrictMono α] [MulPosStrictMono α] - [CovariantClass α α (· + ·) (· ≤ ·)] [ContravariantClass α α (· + ·) (· ≤ ·)] : + [AddLeftMono α] [AddLeftReflectLE α] : a * b ≤ 0 ↔ (a < 0 → 0 ≤ b) ∧ (0 < b → a ≤ 0) := by rw [← neg_nonneg, ← neg_mul, mul_nonneg_iff_pos_imp_nonneg (α := α)] simp only [neg_pos, neg_nonneg] lemma neg_one_lt_zero - [ZeroLEOneClass α] [NeZero (R := α) 1] [CovariantClass α α (· + ·) (· < ·)] : + [ZeroLEOneClass α] [NeZero (R := α) 1] [AddLeftStrictMono α] : -1 < (0 : α) := neg_lt_zero.2 zero_lt_one lemma sub_one_lt [ZeroLEOneClass α] [NeZero (R := α) 1] - [CovariantClass α α (· + ·) (· < ·)] + [AddLeftStrictMono α] (a : α) : a - 1 < a := sub_lt_iff_lt_add.2 <| lt_add_one a lemma mul_self_le_mul_self_of_le_of_neg_le - [MulPosMono α] [PosMulMono α] [CovariantClass α α (· + ·) (· ≤ ·)] + [MulPosMono α] [PosMulMono α] [AddLeftMono α] (h₁ : a ≤ b) (h₂ : -a ≤ b) : a * a ≤ b * b := (le_total 0 a).elim (mul_self_le_mul_self · h₁) fun h ↦ (neg_mul_neg a a).symm.trans_le <| diff --git a/Mathlib/Algebra/Order/Ring/Unbundled/Nonneg.lean b/Mathlib/Algebra/Order/Ring/Unbundled/Nonneg.lean index 672f46dd53d38..1993a125516c0 100644 --- a/Mathlib/Algebra/Order/Ring/Unbundled/Nonneg.lean +++ b/Mathlib/Algebra/Order/Ring/Unbundled/Nonneg.lean @@ -98,32 +98,30 @@ theorem mk_eq_zero [Zero α] [Preorder α] {x : α} (hx : 0 ≤ x) : (⟨x, hx⟩ : { x : α // 0 ≤ x }) = 0 ↔ x = 0 := Subtype.ext_iff -instance add [AddZeroClass α] [Preorder α] [CovariantClass α α (· + ·) (· ≤ ·)] : - Add { x : α // 0 ≤ x } := +instance add [AddZeroClass α] [Preorder α] [AddLeftMono α] : Add { x : α // 0 ≤ x } := ⟨fun x y => ⟨x + y, add_nonneg x.2 y.2⟩⟩ @[simp] -theorem mk_add_mk [AddZeroClass α] [Preorder α] [CovariantClass α α (· + ·) (· ≤ ·)] {x y : α} +theorem mk_add_mk [AddZeroClass α] [Preorder α] [AddLeftMono α] {x y : α} (hx : 0 ≤ x) (hy : 0 ≤ y) : (⟨x, hx⟩ : { x : α // 0 ≤ x }) + ⟨y, hy⟩ = ⟨x + y, add_nonneg hx hy⟩ := rfl @[simp, norm_cast] -protected theorem coe_add [AddZeroClass α] [Preorder α] [CovariantClass α α (· + ·) (· ≤ ·)] +protected theorem coe_add [AddZeroClass α] [Preorder α] [AddLeftMono α] (a b : { x : α // 0 ≤ x }) : ((a + b : { x : α // 0 ≤ x }) : α) = a + b := rfl -instance nsmul [AddMonoid α] [Preorder α] [CovariantClass α α (· + ·) (· ≤ ·)] : - SMul ℕ { x : α // 0 ≤ x } := +instance nsmul [AddMonoid α] [Preorder α] [AddLeftMono α] : SMul ℕ { x : α // 0 ≤ x } := ⟨fun n x => ⟨n • (x : α), nsmul_nonneg x.prop n⟩⟩ @[simp] -theorem nsmul_mk [AddMonoid α] [Preorder α] [CovariantClass α α (· + ·) (· ≤ ·)] (n : ℕ) {x : α} +theorem nsmul_mk [AddMonoid α] [Preorder α] [AddLeftMono α] (n : ℕ) {x : α} (hx : 0 ≤ x) : (n • (⟨x, hx⟩ : { x : α // 0 ≤ x })) = ⟨n • x, nsmul_nonneg hx n⟩ := rfl @[simp, norm_cast] -protected theorem coe_nsmul [AddMonoid α] [Preorder α] [CovariantClass α α (· + ·) (· ≤ ·)] +protected theorem coe_nsmul [AddMonoid α] [Preorder α] [AddLeftMono α] (n : ℕ) (a : { x : α // 0 ≤ x }) : ((n • a : { x : α // 0 ≤ x }) : α) = n • (a : α) := rfl @@ -166,7 +164,7 @@ end Mul section AddMonoid -variable [AddMonoid α] [Preorder α] [CovariantClass α α (· + ·) (· ≤ ·)] +variable [AddMonoid α] [Preorder α] [AddLeftMono α] instance addMonoid : AddMonoid { x : α // 0 ≤ x } := Subtype.coe_injective.addMonoid _ Nonneg.coe_zero (fun _ _ => rfl) fun _ _ => rfl @@ -186,7 +184,7 @@ end AddMonoid section AddCommMonoid -variable [AddCommMonoid α] [Preorder α] [CovariantClass α α (· + ·) (· ≤ ·)] +variable [AddCommMonoid α] [Preorder α] [AddLeftMono α] instance addCommMonoid : AddCommMonoid { x : α // 0 ≤ x } := Subtype.coe_injective.addCommMonoid _ Nonneg.coe_zero (fun _ _ => rfl) (fun _ _ => rfl) @@ -195,8 +193,7 @@ end AddCommMonoid section AddMonoidWithOne -variable [AddMonoidWithOne α] [PartialOrder α] -variable [CovariantClass α α (· + ·) (· ≤ ·)] [ZeroLEOneClass α] +variable [AddMonoidWithOne α] [PartialOrder α] [AddLeftMono α] [ZeroLEOneClass α] instance natCast : NatCast { x : α // 0 ≤ x } := ⟨fun n => ⟨n, Nat.cast_nonneg' n⟩⟩ @@ -254,7 +251,7 @@ end Pow section Semiring variable [Semiring α] [PartialOrder α] [ZeroLEOneClass α] - [CovariantClass α α (· + ·) (· ≤ ·)] [PosMulMono α] + [AddLeftMono α] [PosMulMono α] instance semiring : Semiring { x : α // 0 ≤ x } := Subtype.coe_injective.semiring _ Nonneg.coe_zero Nonneg.coe_one @@ -276,7 +273,7 @@ end Semiring section CommSemiring variable [CommSemiring α] [PartialOrder α] [ZeroLEOneClass α] - [CovariantClass α α (· + ·) (· ≤ ·)] [PosMulMono α] + [AddLeftMono α] [PosMulMono α] instance commSemiring : CommSemiring { x : α // 0 ≤ x } := Subtype.coe_injective.commSemiring _ Nonneg.coe_zero Nonneg.coe_one diff --git a/Mathlib/Algebra/Order/Ring/Unbundled/Rat.lean b/Mathlib/Algebra/Order/Ring/Unbundled/Rat.lean index 857fea4b6cfde..4ad83938866b6 100644 --- a/Mathlib/Algebra/Order/Ring/Unbundled/Rat.lean +++ b/Mathlib/Algebra/Order/Ring/Unbundled/Rat.lean @@ -191,7 +191,7 @@ protected lemma lt_def : p < q ↔ p.num * q.den < q.num * p.den := by protected theorem add_le_add_left {a b c : ℚ} : c + a ≤ c + b ↔ a ≤ b := by rw [Rat.le_iff_sub_nonneg, add_sub_add_left_eq_sub, ← Rat.le_iff_sub_nonneg] -instance : CovariantClass ℚ ℚ (· + ·) (· ≤ ·) where +instance : AddLeftMono ℚ where elim := fun _ _ _ h => Rat.add_le_add_left.2 h @[simp] lemma num_nonpos {a : ℚ} : a.num ≤ 0 ↔ a ≤ 0 := by diff --git a/Mathlib/Algebra/Order/Star/Basic.lean b/Mathlib/Algebra/Order/Star/Basic.lean index 6df890e8638a3..19279611c7f7f 100644 --- a/Mathlib/Algebra/Order/Star/Basic.lean +++ b/Mathlib/Algebra/Order/Star/Basic.lean @@ -112,7 +112,7 @@ lemma of_nonneg_iff [NonUnitalRing R] [PartialOrder R] [StarRing R] (h_nonneg_iff : ∀ x : R, 0 ≤ x ↔ x ∈ AddSubmonoid.closure (Set.range fun s : R => star s * s)) : StarOrderedRing R where le_iff x y := by - haveI : CovariantClass R R (· + ·) (· ≤ ·) := ⟨fun _ _ _ h => h_add h _⟩ + have : AddLeftMono R := ⟨fun _ _ _ h => h_add h _⟩ simpa only [← sub_eq_iff_eq_add', sub_nonneg, exists_eq_right'] using h_nonneg_iff (y - x) /-- When `R` is a non-unital ring, to construct a `StarOrderedRing` instance it suffices to @@ -126,7 +126,7 @@ lemma of_nonneg_iff' [NonUnitalRing R] [PartialOrder R] [StarRing R] (h_add : ∀ {x y : R}, x ≤ y → ∀ z, z + x ≤ z + y) (h_nonneg_iff : ∀ x : R, 0 ≤ x ↔ ∃ s, x = star s * s) : StarOrderedRing R := of_le_iff <| by - haveI : CovariantClass R R (· + ·) (· ≤ ·) := ⟨fun _ _ _ h => h_add h _⟩ + have : AddLeftMono R := ⟨fun _ _ _ h => h_add h _⟩ simpa [sub_eq_iff_eq_add', sub_nonneg] using fun x y => h_nonneg_iff (y - x) theorem nonneg_iff [NonUnitalSemiring R] [PartialOrder R] [StarRing R] [StarOrderedRing R] {x : R} : diff --git a/Mathlib/Algebra/Order/Sub/Basic.lean b/Mathlib/Algebra/Order/Sub/Basic.lean index e1b98ddf85486..2bab771de332d 100644 --- a/Mathlib/Algebra/Order/Sub/Basic.lean +++ b/Mathlib/Algebra/Order/Sub/Basic.lean @@ -74,7 +74,7 @@ end AddLECancellable section Contra -variable [ContravariantClass α α (· + ·) (· ≤ ·)] +variable [AddLeftReflectLE α] theorem tsub_le_tsub_iff_left (h : c ≤ a) : a - b ≤ a - c ↔ c ≤ b := Contravariant.AddLECancellable.tsub_le_tsub_iff_left Contravariant.AddLECancellable h @@ -145,7 +145,7 @@ end AddLECancellable section Contra -variable [ContravariantClass α α (· + ·) (· ≤ ·)] +variable [AddLeftReflectLE α] /-- This lemma also holds for `ENNReal`, but we need a different proof for that. -/ theorem tsub_lt_tsub_iff_right (h : c ≤ a) : a - c < b - c ↔ a < b := @@ -186,8 +186,7 @@ theorem tsub_add_min : a - b + min a b = a := by apply min_le_left -- `Odd.tsub` requires `CanonicallyLinearOrderedSemiring`, which we don't have -lemma Even.tsub - [ContravariantClass α α (· + ·) (· ≤ ·)] {m n : α} (hm : Even m) (hn : Even n) : +lemma Even.tsub [AddLeftReflectLE α] {m n : α} (hm : Even m) (hn : Even n) : Even (m - n) := by obtain ⟨a, rfl⟩ := hm obtain ⟨b, rfl⟩ := hn diff --git a/Mathlib/Algebra/Order/Sub/Defs.lean b/Mathlib/Algebra/Order/Sub/Defs.lean index 9966239cd270f..e4c6643431795 100644 --- a/Mathlib/Algebra/Order/Sub/Defs.lean +++ b/Mathlib/Algebra/Order/Sub/Defs.lean @@ -30,7 +30,7 @@ implications if a bi-implication can be proven under the same assumptions. Lemmas using this class are named using `tsub` instead of `sub` (short for "truncated subtraction"). This is to avoid naming conflicts with similar lemmas about ordered groups. -We provide a second version of most results that require `[ContravariantClass α α (+) (≤)]`. In the +We provide a second version of most results that require `[AddLeftReflectLE α]`. In the second version we replace this type-class assumption by explicit `AddLECancellable` assumptions. TODO: maybe we should make a multiplicative version of this, so that we can replace some identical @@ -62,7 +62,7 @@ theorem tsub_le_iff_right [LE α] [Add α] [Sub α] [OrderedSub α] {a b c : α} variable [Preorder α] [Add α] [Sub α] [OrderedSub α] {a b : α} -/-- See `add_tsub_cancel_right` for the equality if `ContravariantClass α α (+) (≤)`. -/ +/-- See `add_tsub_cancel_right` for the equality if `AddLeftReflectLE α`. -/ theorem add_tsub_le_right : a + b - b ≤ a := tsub_le_iff_right.mpr le_rfl @@ -90,7 +90,7 @@ theorem tsub_le_iff_left : a - b ≤ c ↔ a ≤ b + c := by rw [tsub_le_iff_rig theorem le_add_tsub : a ≤ b + (a - b) := tsub_le_iff_left.mp le_rfl -/-- See `add_tsub_cancel_left` for the equality if `ContravariantClass α α (+) (≤)`. -/ +/-- See `add_tsub_cancel_left` for the equality if `AddLeftReflectLE α`. -/ theorem add_tsub_le_left : a + b - a ≤ b := tsub_le_iff_left.mpr le_rfl @@ -105,7 +105,7 @@ theorem tsub_tsub_le : b - (b - a) ≤ a := section Cov -variable [CovariantClass α α (· + ·) (· ≤ ·)] +variable [AddLeftMono α] @[gcongr] theorem tsub_le_tsub_left (h : a ≤ b) (c : α) : c - b ≤ c - a := tsub_le_iff_left.mpr <| le_add_tsub.trans <| add_le_add_right h _ @@ -192,7 +192,7 @@ end AddLECancellable section Contra -variable [ContravariantClass α α (· + ·) (· ≤ ·)] +variable [AddLeftReflectLE α] theorem le_add_tsub_swap : a ≤ b + a - b := Contravariant.AddLECancellable.le_add_tsub_swap @@ -252,7 +252,7 @@ protected theorem tsub_eq_of_eq_add (hb : AddLECancellable b) (h : a = c + b) : /-- Weaker version of `AddLECancellable.tsub_eq_of_eq_add` assuming that `a = c + b` itself is cancellable rather than `b`. -/ -protected lemma tsub_eq_of_eq_add' [CovariantClass α α (· + ·) (· ≤ ·)] (ha : AddLECancellable a) +protected lemma tsub_eq_of_eq_add' [AddLeftMono α] (ha : AddLECancellable a) (h : a = c + b) : a - b = c := (h ▸ ha).of_add_right.tsub_eq_of_eq_add h /-- See `AddLECancellable.eq_tsub_of_add_eq'` for a version assuming that `b = a + c` itself is @@ -262,7 +262,7 @@ protected theorem eq_tsub_of_add_eq (hc : AddLECancellable c) (h : a + c = b) : /-- Weaker version of `AddLECancellable.eq_tsub_of_add_eq` assuming that `b = a + c` itself is cancellable rather than `c`. -/ -protected lemma eq_tsub_of_add_eq' [CovariantClass α α (· + ·) (· ≤ ·)] (hb : AddLECancellable b) +protected lemma eq_tsub_of_add_eq' [AddLeftMono α] (hb : AddLECancellable b) (h : a + c = b) : a = b - c := (hb.tsub_eq_of_eq_add' h.symm).symm /-- See `AddLECancellable.tsub_eq_of_eq_add_rev'` for a version assuming that `a = b + c` itself is @@ -272,7 +272,7 @@ protected theorem tsub_eq_of_eq_add_rev (hb : AddLECancellable b) (h : a = b + c /-- Weaker version of `AddLECancellable.tsub_eq_of_eq_add_rev` assuming that `a = b + c` itself is cancellable rather than `b`. -/ -protected lemma tsub_eq_of_eq_add_rev' [CovariantClass α α (· + ·) (· ≤ ·)] +protected lemma tsub_eq_of_eq_add_rev' [AddLeftMono α] (ha : AddLECancellable a) (h : a = b + c) : a - b = c := ha.tsub_eq_of_eq_add' <| by rw [add_comm, h] @@ -312,7 +312,7 @@ end AddLECancellable section Contra -variable [ContravariantClass α α (· + ·) (· ≤ ·)] +variable [AddLeftReflectLE α] theorem tsub_eq_of_eq_add (h : a = c + b) : a - b = c := Contravariant.AddLECancellable.tsub_eq_of_eq_add h @@ -355,7 +355,7 @@ end Contra section Both -variable [CovariantClass α α (· + ·) (· ≤ ·)] [ContravariantClass α α (· + ·) (· ≤ ·)] +variable [AddLeftMono α] [AddLeftReflectLE α] theorem add_tsub_add_eq_tsub_right (a c b : α) : a + c - (b + c) = a - b := by refine add_tsub_add_le_tsub_right.antisymm (tsub_le_iff_right.2 <| ?_) @@ -394,7 +394,7 @@ theorem lt_tsub_comm : a < b - c ↔ c < b - a := section Cov -variable [CovariantClass α α (· + ·) (· ≤ ·)] +variable [AddLeftMono α] /-- See `lt_of_tsub_lt_tsub_left_of_le` for a weaker statement in a partial order. -/ theorem lt_of_tsub_lt_tsub_left (h : a - b < a - c) : c < b := diff --git a/Mathlib/Algebra/Order/Sub/Unbundled/Basic.lean b/Mathlib/Algebra/Order/Sub/Unbundled/Basic.lean index 979e4903eb48f..899776822e162 100644 --- a/Mathlib/Algebra/Order/Sub/Unbundled/Basic.lean +++ b/Mathlib/Algebra/Order/Sub/Unbundled/Basic.lean @@ -18,7 +18,7 @@ variable {α : Type*} section ExistsAddOfLE variable [AddCommSemigroup α] [PartialOrder α] [ExistsAddOfLE α] - [CovariantClass α α (· + ·) (· ≤ ·)] [Sub α] [OrderedSub α] {a b c d : α} + [AddLeftMono α] [Sub α] [OrderedSub α] {a b c d : α} @[simp] theorem add_tsub_cancel_of_le (h : a ≤ b) : a + (b - a) = b := by @@ -130,7 +130,7 @@ protected theorem tsub_inj_right (hab : AddLECancellable (a - b)) (h₁ : b ≤ rw [← hab.inj] rw [tsub_add_cancel_of_le h₁, h₃, tsub_add_cancel_of_le h₂] -protected theorem lt_of_tsub_lt_tsub_left_of_le [ContravariantClass α α (· + ·) (· < ·)] +protected theorem lt_of_tsub_lt_tsub_left_of_le [AddLeftReflectLT α] (hb : AddLECancellable b) (hca : c ≤ a) (h : a - b < a - c) : c < b := by conv_lhs at h => rw [← tsub_add_cancel_of_le hca] exact lt_of_add_lt_add_left (hb.lt_add_of_tsub_lt_right h) @@ -144,7 +144,7 @@ protected theorem tsub_lt_tsub_right_of_le (hc : AddLECancellable c) (h : c ≤ apply hc.lt_tsub_of_add_lt_left rwa [add_tsub_cancel_of_le h] -protected theorem tsub_lt_tsub_iff_left_of_le_of_le [ContravariantClass α α (· + ·) (· < ·)] +protected theorem tsub_lt_tsub_iff_left_of_le_of_le [AddLeftReflectLT α] (hb : AddLECancellable b) (hab : AddLECancellable (a - b)) (h₁ : b ≤ a) (h₂ : c ≤ a) : a - b < a - c ↔ c < b := ⟨hb.lt_of_tsub_lt_tsub_left_of_le h₂, hab.tsub_lt_tsub_left_of_le h₁⟩ @@ -168,7 +168,7 @@ section Contra /-! ### Lemmas where addition is order-reflecting. -/ -variable [ContravariantClass α α (· + ·) (· ≤ ·)] +variable [AddLeftReflectLE α] theorem eq_tsub_iff_add_eq_of_le (h : c ≤ b) : a = b - c ↔ a + c = b := Contravariant.AddLECancellable.eq_tsub_iff_add_eq_of_le h @@ -216,7 +216,7 @@ theorem lt_tsub_iff_left_of_le (h : c ≤ b) : a < b - c ↔ c + a < b := Contravariant.AddLECancellable.lt_tsub_iff_left_of_le h /-- See `lt_of_tsub_lt_tsub_left` for a stronger statement in a linear order. -/ -theorem lt_of_tsub_lt_tsub_left_of_le [ContravariantClass α α (· + ·) (· < ·)] (hca : c ≤ a) +theorem lt_of_tsub_lt_tsub_left_of_le [AddLeftReflectLT α] (hca : c ≤ a) (h : a - b < a - c) : c < b := Contravariant.AddLECancellable.lt_of_tsub_lt_tsub_left_of_le hca h @@ -230,7 +230,7 @@ theorem tsub_inj_right (h₁ : b ≤ a) (h₂ : c ≤ a) (h₃ : a - b = a - c) Contravariant.AddLECancellable.tsub_inj_right h₁ h₂ h₃ /-- See `tsub_lt_tsub_iff_left_of_le` for a stronger statement in a linear order. -/ -theorem tsub_lt_tsub_iff_left_of_le_of_le [ContravariantClass α α (· + ·) (· < ·)] (h₁ : b ≤ a) +theorem tsub_lt_tsub_iff_left_of_le_of_le [AddLeftReflectLT α] (h₁ : b ≤ a) (h₂ : c ≤ a) : a - b < a - c ↔ c < b := Contravariant.AddLECancellable.tsub_lt_tsub_iff_left_of_le_of_le Contravariant.AddLECancellable h₁ h₂ diff --git a/Mathlib/Algebra/Order/Sub/Unbundled/Hom.lean b/Mathlib/Algebra/Order/Sub/Unbundled/Hom.lean index fedfe17b0c39c..6d9cedb09fb7c 100644 --- a/Mathlib/Algebra/Order/Sub/Unbundled/Hom.lean +++ b/Mathlib/Algebra/Order/Sub/Unbundled/Hom.lean @@ -24,11 +24,11 @@ theorem AddHom.le_map_tsub [Preorder β] [Add β] [Sub β] [OrderedSub β] (f : exact hf le_tsub_add theorem le_mul_tsub {R : Type*} [Distrib R] [Preorder R] [Sub R] [OrderedSub R] - [CovariantClass R R (· * ·) (· ≤ ·)] {a b c : R} : a * b - a * c ≤ a * (b - c) := + [MulLeftMono R] {a b c : R} : a * b - a * c ≤ a * (b - c) := (AddHom.mulLeft a).le_map_tsub (monotone_id.const_mul' a) _ _ theorem le_tsub_mul {R : Type*} [CommSemiring R] [Preorder R] [Sub R] [OrderedSub R] - [CovariantClass R R (· * ·) (· ≤ ·)] {a b c : R} : a * c - b * c ≤ (a - b) * c := by + [MulLeftMono R] {a b c : R} : a * c - b * c ≤ (a - b) * c := by simpa only [mul_comm _ c] using le_mul_tsub end Add diff --git a/Mathlib/Algebra/Tropical/Basic.lean b/Mathlib/Algebra/Tropical/Basic.lean index 5454363d36711..373542c05d885 100644 --- a/Mathlib/Algebra/Tropical/Basic.lean +++ b/Mathlib/Algebra/Tropical/Basic.lean @@ -431,15 +431,15 @@ end Monoid section Distrib -instance covariant_mul [LE R] [Add R] [CovariantClass R R (· + ·) (· ≤ ·)] : - CovariantClass (Tropical R) (Tropical R) (· * ·) (· ≤ ·) := +instance mulLeftMono [LE R] [Add R] [AddLeftMono R] : + MulLeftMono (Tropical R) := ⟨fun _ y z h => add_le_add_left (show untrop y ≤ untrop z from h) _⟩ -instance covariant_swap_mul [LE R] [Add R] [CovariantClass R R (Function.swap (· + ·)) (· ≤ ·)] : - CovariantClass (Tropical R) (Tropical R) (Function.swap (· * ·)) (· ≤ ·) := +instance mulRightMono [LE R] [Add R] [AddRightMono R] : + MulRightMono (Tropical R) := ⟨fun _ y z h => add_le_add_right (show untrop y ≤ untrop z from h) _⟩ -instance covariant_add [LinearOrder R] : CovariantClass (Tropical R) (Tropical R) (· + ·) (· ≤ ·) := +instance addLeftMono [LinearOrder R] : AddLeftMono (Tropical R) := ⟨fun x y z h => by rcases le_total x y with hx | hy · rw [add_eq_left hx, add_eq_left (hx.trans h)] @@ -448,17 +448,15 @@ instance covariant_add [LinearOrder R] : CovariantClass (Tropical R) (Tropical R · rwa [add_eq_left hx] · rwa [add_eq_right hx]⟩ -instance covariant_mul_lt [LT R] [Add R] [CovariantClass R R (· + ·) (· < ·)] : - CovariantClass (Tropical R) (Tropical R) (· * ·) (· < ·) := +instance mulLeftStrictMono [LT R] [Add R] [AddLeftStrictMono R] : + MulLeftStrictMono (Tropical R) := ⟨fun _ _ _ h => add_lt_add_left (untrop_lt_iff.2 h) _⟩ -instance covariant_swap_mul_lt [Preorder R] [Add R] - [CovariantClass R R (Function.swap (· + ·)) (· < ·)] : - CovariantClass (Tropical R) (Tropical R) (Function.swap (· * ·)) (· < ·) := +instance mulRightStrictMono [Preorder R] [Add R] [AddRightStrictMono R] : + MulRightStrictMono (Tropical R) := ⟨fun _ y z h => add_lt_add_right (show untrop y < untrop z from h) _⟩ -instance instDistribTropical [LinearOrder R] [Add R] [CovariantClass R R (· + ·) (· ≤ ·)] - [CovariantClass R R (Function.swap (· + ·)) (· ≤ ·)] : +instance instDistribTropical [LinearOrder R] [Add R] [AddLeftMono R] [AddRightMono R] : Distrib (Tropical R) where mul := (· * ·) add := (· + ·) @@ -466,8 +464,8 @@ instance instDistribTropical [LinearOrder R] [Add R] [CovariantClass R R (· + right_distrib _ _ _ := untrop_injective (min_add_add_right _ _ _).symm @[simp] -theorem add_pow [LinearOrder R] [AddMonoid R] [CovariantClass R R (· + ·) (· ≤ ·)] - [CovariantClass R R (Function.swap (· + ·)) (· ≤ ·)] (x y : Tropical R) (n : ℕ) : +theorem add_pow [LinearOrder R] [AddMonoid R] [AddLeftMono R] [AddRightMono R] + (x y : Tropical R) (n : ℕ) : (x + y) ^ n = x ^ n + y ^ n := by rcases le_total x y with h | h · rw [add_eq_left h, add_eq_left (pow_le_pow_left' h _)] diff --git a/Mathlib/Combinatorics/SetFamily/CauchyDavenport.lean b/Mathlib/Combinatorics/SetFamily/CauchyDavenport.lean index a08ee82af2783..9bd497ebf7c8b 100644 --- a/Mathlib/Combinatorics/SetFamily/CauchyDavenport.lean +++ b/Mathlib/Combinatorics/SetFamily/CauchyDavenport.lean @@ -202,7 +202,7 @@ lemma ZMod.min_le_card_add {p : ℕ} (hp : p.Prime) {s t : Finset (ZMod p)} (hs "The **Cauchy-Davenport theorem** for linearly ordered additive cancellative semigroups. The size of `s + t` is lower-bounded by `|s| + |t| - 1`."] lemma Finset.card_add_card_sub_one_le_card_mul [LinearOrder α] [Semigroup α] [IsCancelMul α] - [CovariantClass α α (· * ·) (· ≤ ·)] [CovariantClass α α (swap (· * ·)) (· ≤ ·)] + [MulLeftMono α] [MulRightMono α] {s t : Finset α} (hs : s.Nonempty) (ht : t.Nonempty) : #s + #t - 1 ≤ #(s * t) := by suffices s * {t.min' ht} ∩ ({s.max' hs} * t) = {s.max' hs * t.min' ht} by rw [← card_singleton_mul t (s.max' hs), ← card_mul_singleton s (t.min' ht), diff --git a/Mathlib/Data/DFinsupp/Lex.lean b/Mathlib/Data/DFinsupp/Lex.lean index b0c8d8ea0b694..1b3888f712e29 100644 --- a/Mathlib/Data/DFinsupp/Lex.lean +++ b/Mathlib/Data/DFinsupp/Lex.lean @@ -136,37 +136,33 @@ section Covariants variable [LinearOrder ι] [∀ i, AddMonoid (α i)] [∀ i, LinearOrder (α i)] /-! We are about to sneak in a hypothesis that might appear to be too strong. -We assume `CovariantClass` with *strict* inequality `<` also when proving the one with the -*weak* inequality `≤`. This is actually necessary: addition on `Lex (Π₀ i, α i)` may fail to be -monotone, when it is "just" monotone on `α i`. -/ +We assume `AddLeftStrictMono` (covariant with *strict* inequality `<`) also when proving the one +with the *weak* inequality `≤`. This is actually necessary: addition on `Lex (Π₀ i, α i)` may fail +to be monotone, when it is "just" monotone on `α i`. -/ section Left -variable [∀ i, CovariantClass (α i) (α i) (· + ·) (· < ·)] +variable [∀ i, AddLeftStrictMono (α i)] -instance Lex.covariantClass_lt_left : - CovariantClass (Lex (Π₀ i, α i)) (Lex (Π₀ i, α i)) (· + ·) (· < ·) := +instance Lex.addLeftStrictMono : AddLeftStrictMono (Lex (Π₀ i, α i)) := ⟨fun _ _ _ ⟨a, lta, ha⟩ ↦ ⟨a, fun j ja ↦ congr_arg _ (lta j ja), add_lt_add_left ha _⟩⟩ -instance Lex.covariantClass_le_left : - CovariantClass (Lex (Π₀ i, α i)) (Lex (Π₀ i, α i)) (· + ·) (· ≤ ·) := - covariantClass_le_of_lt _ _ _ +instance Lex.addLeftMono : AddLeftMono (Lex (Π₀ i, α i)) := + addLeftMono_of_addLeftStrictMono _ end Left section Right -variable [∀ i, CovariantClass (α i) (α i) (Function.swap (· + ·)) (· < ·)] +variable [∀ i, AddRightStrictMono (α i)] -instance Lex.covariantClass_lt_right : - CovariantClass (Lex (Π₀ i, α i)) (Lex (Π₀ i, α i)) (Function.swap (· + ·)) (· < ·) := +instance Lex.addRightStrictMono : AddRightStrictMono (Lex (Π₀ i, α i)) := ⟨fun f _ _ ⟨a, lta, ha⟩ ↦ ⟨a, fun j ja ↦ congr_arg (· + ofLex f j) (lta j ja), add_lt_add_right ha _⟩⟩ -instance Lex.covariantClass_le_right : - CovariantClass (Lex (Π₀ i, α i)) (Lex (Π₀ i, α i)) (Function.swap (· + ·)) (· ≤ ·) := - covariantClass_le_of_lt _ _ _ +instance Lex.addRightMono : AddRightMono (Lex (Π₀ i, α i)) := + addRightMono_of_addRightStrictMono _ end Right diff --git a/Mathlib/Data/DFinsupp/Order.lean b/Mathlib/Data/DFinsupp/Order.lean index 6ace1dada4df5..767a61edb7f6b 100644 --- a/Mathlib/Data/DFinsupp/Order.lean +++ b/Mathlib/Data/DFinsupp/Order.lean @@ -141,8 +141,8 @@ instance (α : ι → Type*) [∀ i, OrderedCancelAddCommMonoid (α i)] : { (inferInstance : OrderedAddCommMonoid (DFinsupp α)) with le_of_add_le_add_left := fun _ _ _ H i ↦ le_of_add_le_add_left (H i) } -instance [∀ i, OrderedAddCommMonoid (α i)] [∀ i, ContravariantClass (α i) (α i) (· + ·) (· ≤ ·)] : - ContravariantClass (Π₀ i, α i) (Π₀ i, α i) (· + ·) (· ≤ ·) := +instance [∀ i, OrderedAddCommMonoid (α i)] [∀ i, AddLeftReflectLE (α i)] : + AddLeftReflectLE (Π₀ i, α i) := ⟨fun _ _ _ H i ↦ le_of_add_le_add_left (H i)⟩ section Module diff --git a/Mathlib/Data/ENNReal/Basic.lean b/Mathlib/Data/ENNReal/Basic.lean index 737f546515208..6aac4f856218f 100644 --- a/Mathlib/Data/ENNReal/Basic.lean +++ b/Mathlib/Data/ENNReal/Basic.lean @@ -132,9 +132,9 @@ noncomputable instance : DivInvMonoid ℝ≥0∞ where variable {a b c d : ℝ≥0∞} {r p q : ℝ≥0} -- Porting note: are these 2 instances still required in Lean 4? -instance covariantClass_mul_le : CovariantClass ℝ≥0∞ ℝ≥0∞ (· * ·) (· ≤ ·) := inferInstance +instance mulLeftMono : MulLeftMono ℝ≥0∞ := inferInstance -instance covariantClass_add_le : CovariantClass ℝ≥0∞ ℝ≥0∞ (· + ·) (· ≤ ·) := inferInstance +instance addLeftMono : AddLeftMono ℝ≥0∞ := inferInstance -- Porting note (#11215): TODO: add a `WithTop` instance and use it here noncomputable instance : LinearOrderedCommMonoidWithZero ℝ≥0∞ := diff --git a/Mathlib/Data/ENNReal/Operations.lean b/Mathlib/Data/ENNReal/Operations.lean index 6931936b8b2fe..55e43c9ea0a79 100644 --- a/Mathlib/Data/ENNReal/Operations.lean +++ b/Mathlib/Data/ENNReal/Operations.lean @@ -39,10 +39,10 @@ theorem mul_lt_mul (ac : a < c) (bd : b < d) : a * b < c * d := by ↑(a * b) < ↑(a' * b') := coe_lt_coe.2 (mul_lt_mul₀ aa' bb') _ ≤ c * d := mul_le_mul' a'c.le b'd.le --- TODO: generalize to `CovariantClass α α (· * ·) (· ≤ ·)` +-- TODO: generalize to `MulLeftMono α` theorem mul_left_mono : Monotone (a * ·) := fun _ _ => mul_le_mul' le_rfl --- TODO: generalize to `CovariantClass α α (swap (· * ·)) (· ≤ ·)` +-- TODO: generalize to `MulRightMono α` theorem mul_right_mono : Monotone (· * a) := fun _ _ h => mul_le_mul' h le_rfl -- Porting note (#11215): TODO: generalize to `WithTop` @@ -143,8 +143,8 @@ protected theorem add_lt_add_of_le_of_lt : a ≠ ∞ → a ≤ b → c < d → a protected theorem add_lt_add_of_lt_of_le : c ≠ ∞ → a < b → c ≤ d → a + c < b + d := WithTop.add_lt_add_of_lt_of_le -instance contravariantClass_add_lt : ContravariantClass ℝ≥0∞ ℝ≥0∞ (· + ·) (· < ·) := - WithTop.contravariantClass_add_lt +instance addLeftReflectLT : AddLeftReflectLT ℝ≥0∞ := + WithTop.addLeftReflectLT theorem lt_add_right (ha : a ≠ ∞) (hb : b ≠ 0) : a < a + b := by rwa [← pos_iff_ne_zero, ← ENNReal.add_lt_add_iff_left ha, add_zero] at hb diff --git a/Mathlib/Data/Finset/Fold.lean b/Mathlib/Data/Finset/Fold.lean index 39befbf6c8239..931d4b86e5a6b 100644 --- a/Mathlib/Data/Finset/Fold.lean +++ b/Mathlib/Data/Finset/Fold.lean @@ -231,7 +231,7 @@ theorem fold_max_lt : s.fold max b f < c ↔ b < c ∧ ∀ x ∈ s, f x < c := b theorem lt_fold_max : c < s.fold max b f ↔ c < b ∨ ∃ x ∈ s, c < f x := fold_op_rel_iff_or lt_max_iff -theorem fold_max_add [Add β] [CovariantClass β β (Function.swap (· + ·)) (· ≤ ·)] (n : WithBot β) +theorem fold_max_add [Add β] [AddRightMono β] (n : WithBot β) (s : Finset α) : (s.fold max ⊥ fun x : α => ↑(f x) + n) = s.fold max ⊥ ((↑) ∘ f) + n := by classical induction' s using Finset.induction_on with a s _ ih <;> simp [*, max_add_add_right] diff --git a/Mathlib/Data/Finset/Lattice.lean b/Mathlib/Data/Finset/Lattice.lean index 6f66f021e480f..ad15a3cd7ffe6 100644 --- a/Mathlib/Data/Finset/Lattice.lean +++ b/Mathlib/Data/Finset/Lattice.lean @@ -814,7 +814,7 @@ theorem _root_.map_finset_sup' [SemilatticeSup β] [FunLike F α β] [SupHomClas refine hs.cons_induction ?_ ?_ <;> intros <;> simp [*] lemma nsmul_sup' {α β : Type*} [AddMonoid β] [LinearOrder β] - [CovariantClass β β (· + ·) (· ≤ ·)] [CovariantClass β β (swap (· + ·)) (· ≤ ·)] + [AddLeftMono β] [AddRightMono β] {s : Finset α} (hs : s.Nonempty) (f : α → β) (n : ℕ) : s.sup' hs (fun a => n • f a) = n • s.sup' hs f := let ns : SupHom β β := { toFun := (n • ·), map_sup' := fun _ _ => (nsmul_right_mono n).map_max } @@ -966,7 +966,7 @@ theorem _root_.map_finset_inf' [SemilatticeInf β] [FunLike F α β] [InfHomClas refine hs.cons_induction ?_ ?_ <;> intros <;> simp [*] lemma nsmul_inf' {α β : Type*} [AddMonoid β] [LinearOrder β] - [CovariantClass β β (· + ·) (· ≤ ·)] [CovariantClass β β (swap (· + ·)) (· ≤ ·)] + [AddLeftMono β] [AddRightMono β] {s : Finset α} (hs : s.Nonempty) (f : α → β) (n : ℕ) : s.inf' hs (fun a => n • f a) = n • s.inf' hs f := let ns : InfHom β β := { toFun := (n • ·), map_inf' := fun _ _ => (nsmul_right_mono n).map_min } diff --git a/Mathlib/Data/Finsupp/Lex.lean b/Mathlib/Data/Finsupp/Lex.lean index 9174fd421d429..6b2a45dd8483b 100644 --- a/Mathlib/Data/Finsupp/Lex.lean +++ b/Mathlib/Data/Finsupp/Lex.lean @@ -91,39 +91,35 @@ section Covariants variable [LinearOrder α] [AddMonoid N] [LinearOrder N] /-! We are about to sneak in a hypothesis that might appear to be too strong. -We assume `CovariantClass` with *strict* inequality `<` also when proving the one with the -*weak* inequality `≤`. This is actually necessary: addition on `Lex (α →₀ N)` may fail to be -monotone, when it is "just" monotone on `N`. +We assume `AddLeftStrictMono` (covariant with *strict* inequality `<`) also when proving the one +with the *weak* inequality `≤`. This is actually necessary: addition on `Lex (α →₀ N)` may fail to +be monotone, when it is "just" monotone on `N`. See `Counterexamples/ZeroDivisorsInAddMonoidAlgebras.lean` for a counterexample. -/ section Left -variable [CovariantClass N N (· + ·) (· < ·)] +variable [AddLeftStrictMono N] -instance Lex.covariantClass_lt_left : - CovariantClass (Lex (α →₀ N)) (Lex (α →₀ N)) (· + ·) (· < ·) := +instance Lex.addLeftStrictMono : AddLeftStrictMono (Lex (α →₀ N)) := ⟨fun _ _ _ ⟨a, lta, ha⟩ ↦ ⟨a, fun j ja ↦ congr_arg _ (lta j ja), add_lt_add_left ha _⟩⟩ -instance Lex.covariantClass_le_left : - CovariantClass (Lex (α →₀ N)) (Lex (α →₀ N)) (· + ·) (· ≤ ·) := - covariantClass_le_of_lt _ _ _ +instance Lex.addLeftMono : AddLeftMono (Lex (α →₀ N)) := + addLeftMono_of_addLeftStrictMono _ end Left section Right -variable [CovariantClass N N (Function.swap (· + ·)) (· < ·)] +variable [AddRightStrictMono N] -instance Lex.covariantClass_lt_right : - CovariantClass (Lex (α →₀ N)) (Lex (α →₀ N)) (Function.swap (· + ·)) (· < ·) := +instance Lex.addRightStrictMono : AddRightStrictMono (Lex (α →₀ N)) := ⟨fun f _ _ ⟨a, lta, ha⟩ ↦ ⟨a, fun j ja ↦ congr_arg (· + ofLex f j) (lta j ja), add_lt_add_right ha _⟩⟩ -instance Lex.covariantClass_le_right : - CovariantClass (Lex (α →₀ N)) (Lex (α →₀ N)) (Function.swap (· + ·)) (· ≤ ·) := - covariantClass_le_of_lt _ _ _ +instance Lex.addRightMono : AddRightMono (Lex (α →₀ N)) := + addRightMono_of_addRightStrictMono _ end Right diff --git a/Mathlib/Data/Finsupp/Order.lean b/Mathlib/Data/Finsupp/Order.lean index cb0d481298010..e1bb9c61c445f 100644 --- a/Mathlib/Data/Finsupp/Order.lean +++ b/Mathlib/Data/Finsupp/Order.lean @@ -171,8 +171,8 @@ instance orderedCancelAddCommMonoid [OrderedCancelAddCommMonoid α] : { Finsupp.orderedAddCommMonoid with le_of_add_le_add_left := fun _f _g _i h s => le_of_add_le_add_left (h s) } -instance contravariantClass [OrderedAddCommMonoid α] [ContravariantClass α α (· + ·) (· ≤ ·)] : - ContravariantClass (ι →₀ α) (ι →₀ α) (· + ·) (· ≤ ·) := +instance addLeftReflectLE [OrderedAddCommMonoid α] [AddLeftReflectLE α] : + AddLeftReflectLE (ι →₀ α) := ⟨fun _f _g _h H x => le_of_add_le_add_left <| H x⟩ section SMulZeroClass diff --git a/Mathlib/Data/Multiset/Basic.lean b/Mathlib/Data/Multiset/Basic.lean index 001ad7a8b6abd..667c2e2f5f6f1 100644 --- a/Mathlib/Data/Multiset/Basic.lean +++ b/Mathlib/Data/Multiset/Basic.lean @@ -588,10 +588,10 @@ theorem singleton_add (a : α) (s : Multiset α) : {a} + s = a ::ₘ s := private theorem add_le_add_iff_left' {s t u : Multiset α} : s + t ≤ s + u ↔ t ≤ u := Quotient.inductionOn₃ s t u fun _ _ _ => subperm_append_left _ -instance : CovariantClass (Multiset α) (Multiset α) (· + ·) (· ≤ ·) := +instance : AddLeftMono (Multiset α) := ⟨fun _s _t _u => add_le_add_iff_left'.2⟩ -instance : ContravariantClass (Multiset α) (Multiset α) (· + ·) (· ≤ ·) := +instance : AddLeftReflectLE (Multiset α) := ⟨fun _s _t _u => add_le_add_iff_left'.1⟩ instance instAddCommMonoid : AddCancelCommMonoid (Multiset α) where diff --git a/Mathlib/Data/NNReal/Defs.lean b/Mathlib/Data/NNReal/Defs.lean index b220a37efdd08..9a80187ab708f 100644 --- a/Mathlib/Data/NNReal/Defs.lean +++ b/Mathlib/Data/NNReal/Defs.lean @@ -440,9 +440,9 @@ theorem coe_iInf {ι : Sort*} (s : ι → ℝ≥0) : (↑(⨅ i, s i) : ℝ) = rw [iInf, iInf, coe_sInf, ← Set.range_comp]; rfl -- Short-circuit instance search -instance instCovariantClassAddLE : CovariantClass ℝ≥0 ℝ≥0 (· + ·) (· ≤ ·) := inferInstance -instance instContravariantClassAddLT : ContravariantClass ℝ≥0 ℝ≥0 (· + ·) (· < ·) := inferInstance -instance instCovariantClassMulLE : CovariantClass ℝ≥0 ℝ≥0 (· * ·) (· ≤ ·) := inferInstance +instance addLeftMono : AddLeftMono ℝ≥0 := inferInstance +instance addLeftReflectLT : AddLeftReflectLT ℝ≥0 := inferInstance +instance mulLeftMono : MulLeftMono ℝ≥0 := inferInstance @[deprecated le_of_forall_pos_le_add (since := "2024-10-17")] protected theorem le_of_forall_pos_le_add {a b : ℝ≥0} (h : ∀ ε, 0 < ε → a ≤ b + ε) : a ≤ b := diff --git a/Mathlib/Data/Nat/Cast/Order/Basic.lean b/Mathlib/Data/Nat/Cast/Order/Basic.lean index 06c79cba42b28..27a78dfb1c92d 100644 --- a/Mathlib/Data/Nat/Cast/Order/Basic.lean +++ b/Mathlib/Data/Nat/Cast/Order/Basic.lean @@ -26,7 +26,7 @@ we use a generic collection of instances so that it applies in other settings (e `StarOrderedRing`, or the `selfAdjoint` or `StarOrderedRing.positive` parts thereof). -/ variable [AddMonoidWithOne α] [PartialOrder α] -variable [CovariantClass α α (· + ·) (· ≤ ·)] [ZeroLEOneClass α] +variable [AddLeftMono α] [ZeroLEOneClass α] @[mono] theorem mono_cast : Monotone (Nat.cast : ℕ → α) := diff --git a/Mathlib/Data/Nat/Cast/Order/Ring.lean b/Mathlib/Data/Nat/Cast/Order/Ring.lean index 2f6687f56c7ac..97687710d7557 100644 --- a/Mathlib/Data/Nat/Cast/Order/Ring.lean +++ b/Mathlib/Data/Nat/Cast/Order/Ring.lean @@ -22,7 +22,7 @@ we use a generic collection of instances so that it applies in other settings (e `StarOrderedRing`, or the `selfAdjoint` or `StarOrderedRing.positive` parts thereof). -/ variable [AddMonoidWithOne α] [PartialOrder α] -variable [CovariantClass α α (· + ·) (· ≤ ·)] [ZeroLEOneClass α] +variable [AddLeftMono α] [ZeroLEOneClass α] /-- Specialisation of `Nat.cast_nonneg'`, which seems to be easier for Lean to use. -/ @[simp] @@ -73,7 +73,7 @@ end OrderedSemiring for `ℕ∞` and `ℝ≥0∞`, so we use type-specific lemmas for these types. -/ @[simp, norm_cast] theorem cast_tsub [CanonicallyOrderedCommSemiring α] [Sub α] [OrderedSub α] - [ContravariantClass α α (· + ·) (· ≤ ·)] (m n : ℕ) : ↑(m - n) = (m - n : α) := by + [AddLeftReflectLE α] (m n : ℕ) : ↑(m - n) = (m - n : α) := by rcases le_total m n with h | h · rw [Nat.sub_eq_zero_of_le h, cast_zero, tsub_eq_zero_of_le] exact mono_cast h diff --git a/Mathlib/Data/PNat/Basic.lean b/Mathlib/Data/PNat/Basic.lean index 4f8bb3bc6f0fc..05f07c00bfc81 100644 --- a/Mathlib/Data/PNat/Basic.lean +++ b/Mathlib/Data/PNat/Basic.lean @@ -117,17 +117,17 @@ def coeAddHom : AddHom ℕ+ ℕ where toFun := Coe.coe map_add' := add_coe -instance covariantClass_add_le : CovariantClass ℕ+ ℕ+ (· + ·) (· ≤ ·) := - Positive.covariantClass_add_le +instance addLeftMono : AddLeftMono ℕ+ := + Positive.addLeftMono -instance covariantClass_add_lt : CovariantClass ℕ+ ℕ+ (· + ·) (· < ·) := - Positive.covariantClass_add_lt +instance addLeftStrictMono : AddLeftStrictMono ℕ+ := + Positive.addLeftStrictMono -instance contravariantClass_add_le : ContravariantClass ℕ+ ℕ+ (· + ·) (· ≤ ·) := - Positive.contravariantClass_add_le +instance addLeftReflectLE : AddLeftReflectLE ℕ+ := + Positive.addLeftReflectLE -instance contravariantClass_add_lt : ContravariantClass ℕ+ ℕ+ (· + ·) (· < ·) := - Positive.contravariantClass_add_lt +instance addLeftReflectLT : AddLeftReflectLT ℕ+ := + Positive.addLeftReflectLT /-- The order isomorphism between ℕ and ℕ+ given by `succ`. -/ @[simps! (config := .asFn) apply] diff --git a/Mathlib/Data/Set/MulAntidiagonal.lean b/Mathlib/Data/Set/MulAntidiagonal.lean index 487d0d5850eb3..8bddb878d7cc2 100644 --- a/Mathlib/Data/Set/MulAntidiagonal.lean +++ b/Mathlib/Data/Set/MulAntidiagonal.lean @@ -83,8 +83,7 @@ end CancelCommMonoid section OrderedCancelCommMonoid -variable [CancelCommMonoid α] [PartialOrder α] [CovariantClass α α (· * ·) (· ≤ ·)] - [CovariantClass α α (Function.swap (· * ·)) (· < ·)] +variable [CancelCommMonoid α] [PartialOrder α] [MulLeftMono α] [MulRightStrictMono α] (s t : Set α) (a : α) {x y : mulAntidiagonal s t a} @[to_additive Set.AddAntidiagonal.eq_of_fst_le_fst_of_snd_le_snd] @@ -112,8 +111,7 @@ theorem finite_of_isPWO (hs : s.IsPWO) (ht : t.IsPWO) (a) : (mulAntidiagonal s t end OrderedCancelCommMonoid -variable [CancelCommMonoid α] [LinearOrder α] [CovariantClass α α (· * ·) (· ≤ ·)] - [CovariantClass α α (Function.swap (· * ·)) (· < ·)] +variable [CancelCommMonoid α] [LinearOrder α] [MulLeftMono α] [MulRightStrictMono α] @[to_additive Set.AddAntidiagonal.finite_of_isWF] theorem finite_of_isWF {s t : Set α} (hs : s.IsWF) (ht : t.IsWF) diff --git a/Mathlib/Data/Set/Pointwise/Interval.lean b/Mathlib/Data/Set/Pointwise/Interval.lean index f1a9b9d837cec..535b6e9100c34 100644 --- a/Mathlib/Data/Set/Pointwise/Interval.lean +++ b/Mathlib/Data/Set/Pointwise/Interval.lean @@ -37,8 +37,7 @@ the unprimed names have been reserved for section ContravariantLE -variable [Mul α] [Preorder α] -variable [CovariantClass α α (· * ·) (· ≤ ·)] [CovariantClass α α (Function.swap HMul.hMul) LE.le] +variable [Mul α] [Preorder α] [MulLeftMono α] [MulRightMono α] @[to_additive Icc_add_Icc_subset] theorem Icc_mul_Icc_subset' (a b c d : α) : Icc a b * Icc c d ⊆ Icc (a * c) (b * d) := by @@ -59,54 +58,57 @@ end ContravariantLE section ContravariantLT -variable [Mul α] [PartialOrder α] -variable [CovariantClass α α (· * ·) (· < ·)] [CovariantClass α α (Function.swap HMul.hMul) LT.lt] +variable [Mul α] [PartialOrder α] [MulLeftStrictMono α] [MulRightStrictMono α] @[to_additive Icc_add_Ico_subset] theorem Icc_mul_Ico_subset' (a b c d : α) : Icc a b * Ico c d ⊆ Ico (a * c) (b * d) := by - haveI := covariantClass_le_of_lt + have := mulLeftMono_of_mulLeftStrictMono α + have := mulRightMono_of_mulRightStrictMono α rintro x ⟨y, ⟨hya, hyb⟩, z, ⟨hzc, hzd⟩, rfl⟩ exact ⟨mul_le_mul' hya hzc, mul_lt_mul_of_le_of_lt hyb hzd⟩ @[to_additive Ico_add_Icc_subset] theorem Ico_mul_Icc_subset' (a b c d : α) : Ico a b * Icc c d ⊆ Ico (a * c) (b * d) := by - haveI := covariantClass_le_of_lt + have := mulLeftMono_of_mulLeftStrictMono α + have := mulRightMono_of_mulRightStrictMono α rintro x ⟨y, ⟨hya, hyb⟩, z, ⟨hzc, hzd⟩, rfl⟩ exact ⟨mul_le_mul' hya hzc, mul_lt_mul_of_lt_of_le hyb hzd⟩ @[to_additive Ioc_add_Ico_subset] theorem Ioc_mul_Ico_subset' (a b c d : α) : Ioc a b * Ico c d ⊆ Ioo (a * c) (b * d) := by - haveI := covariantClass_le_of_lt + have := mulLeftMono_of_mulLeftStrictMono α + have := mulRightMono_of_mulRightStrictMono α rintro x ⟨y, ⟨hya, hyb⟩, z, ⟨hzc, hzd⟩, rfl⟩ exact ⟨mul_lt_mul_of_lt_of_le hya hzc, mul_lt_mul_of_le_of_lt hyb hzd⟩ @[to_additive Ico_add_Ioc_subset] theorem Ico_mul_Ioc_subset' (a b c d : α) : Ico a b * Ioc c d ⊆ Ioo (a * c) (b * d) := by - haveI := covariantClass_le_of_lt + have := mulLeftMono_of_mulLeftStrictMono α + have := mulRightMono_of_mulRightStrictMono α rintro x ⟨y, ⟨hya, hyb⟩, z, ⟨hzc, hzd⟩, rfl⟩ exact ⟨mul_lt_mul_of_le_of_lt hya hzc, mul_lt_mul_of_lt_of_le hyb hzd⟩ @[to_additive Iic_add_Iio_subset] theorem Iic_mul_Iio_subset' (a b : α) : Iic a * Iio b ⊆ Iio (a * b) := by - haveI := covariantClass_le_of_lt + have := mulRightMono_of_mulRightStrictMono α rintro x ⟨y, hya, z, hzb, rfl⟩ exact mul_lt_mul_of_le_of_lt hya hzb @[to_additive Iio_add_Iic_subset] theorem Iio_mul_Iic_subset' (a b : α) : Iio a * Iic b ⊆ Iio (a * b) := by - haveI := covariantClass_le_of_lt + have := mulLeftMono_of_mulLeftStrictMono α rintro x ⟨y, hya, z, hzb, rfl⟩ exact mul_lt_mul_of_lt_of_le hya hzb @[to_additive Ioi_add_Ici_subset] theorem Ioi_mul_Ici_subset' (a b : α) : Ioi a * Ici b ⊆ Ioi (a * b) := by - haveI := covariantClass_le_of_lt + have := mulLeftMono_of_mulLeftStrictMono α rintro x ⟨y, hya, z, hzb, rfl⟩ exact mul_lt_mul_of_lt_of_le hya hzb @[to_additive Ici_add_Ioi_subset] theorem Ici_mul_Ioi_subset' (a b : α) : Ici a * Ioi b ⊆ Ioi (a * b) := by - haveI := covariantClass_le_of_lt + have := mulRightMono_of_mulRightStrictMono α rintro x ⟨y, hya, z, hzb, rfl⟩ exact mul_lt_mul_of_le_of_lt hya hzb diff --git a/Mathlib/Data/Set/Semiring.lean b/Mathlib/Data/Set/Semiring.lean index e6ff5cf305201..e275628926e7c 100644 --- a/Mathlib/Data/Set/Semiring.lean +++ b/Mathlib/Data/Set/Semiring.lean @@ -115,8 +115,8 @@ theorem _root_.Set.up_union (s t : Set α) : up (s ∪ t) = up s + up t := rfl /- Since addition on `SetSemiring` is commutative (it is set union), there is no need -to also have the instance `CovariantClass (SetSemiring α) (SetSemiring α) (swap (+)) (≤)`. -/ -instance covariantClass_add : CovariantClass (SetSemiring α) (SetSemiring α) (· + ·) (· ≤ ·) := +to also have the instance `AddRightMono (SetSemiring α)`. -/ +instance addLeftMono : AddLeftMono (SetSemiring α) := ⟨fun _ _ _ => union_subset_union_right _⟩ section Mul @@ -150,12 +150,10 @@ instance : NoZeroDivisors (SetSemiring α) := b.eq_empty_or_nonempty.resolve_right fun hb => Nonempty.ne_empty ⟨_, mul_mem_mul ha.some_mem hb.some_mem⟩ ab⟩ -instance covariantClass_mul_left : - CovariantClass (SetSemiring α) (SetSemiring α) (· * ·) (· ≤ ·) := +instance mulLeftMono : MulLeftMono (SetSemiring α) := ⟨fun _ _ _ => mul_subset_mul_left⟩ -instance covariantClass_mul_right : - CovariantClass (SetSemiring α) (SetSemiring α) (swap (· * ·)) (· ≤ ·) := +instance mulRightMono : MulRightMono (SetSemiring α) := ⟨fun _ _ _ => mul_subset_mul_right⟩ end Mul diff --git a/Mathlib/Data/Sign.lean b/Mathlib/Data/Sign.lean index 0911620f20e6a..85c7fb46a38a3 100644 --- a/Mathlib/Data/Sign.lean +++ b/Mathlib/Data/Sign.lean @@ -431,7 +431,7 @@ section AddGroup variable [AddGroup α] [Preorder α] [DecidableRel ((· < ·) : α → α → Prop)] -theorem Left.sign_neg [CovariantClass α α (· + ·) (· < ·)] (a : α) : sign (-a) = -sign a := by +theorem Left.sign_neg [AddLeftStrictMono α] (a : α) : sign (-a) = -sign a := by simp_rw [sign_apply, Left.neg_pos_iff, Left.neg_neg_iff] split_ifs with h h' · exact False.elim (lt_asymm h h') @@ -439,7 +439,7 @@ theorem Left.sign_neg [CovariantClass α α (· + ·) (· < ·)] (a : α) : sign · simp · simp -theorem Right.sign_neg [CovariantClass α α (Function.swap (· + ·)) (· < ·)] (a : α) : +theorem Right.sign_neg [AddRightStrictMono α] (a : α) : sign (-a) = -sign a := by simp_rw [sign_apply, Right.neg_pos_iff, Right.neg_neg_iff] split_ifs with h h' diff --git a/Mathlib/LinearAlgebra/Matrix/PosDef.lean b/Mathlib/LinearAlgebra/Matrix/PosDef.lean index 9343bc8070623..0f7ca95450e7f 100644 --- a/Mathlib/LinearAlgebra/Matrix/PosDef.lean +++ b/Mathlib/LinearAlgebra/Matrix/PosDef.lean @@ -150,7 +150,7 @@ protected lemma zpow [StarOrderedRing R] [DecidableEq n] · simpa using hM.pow n · simpa using (hM.pow n).inv -protected lemma add [CovariantClass R R (· + ·) (· ≤ · )] {A : Matrix m m R} {B : Matrix m m R} +protected lemma add [AddLeftMono R] {A : Matrix m m R} {B : Matrix m m R} (hA : A.PosSemidef) (hB : B.PosSemidef) : (A + B).PosSemidef := ⟨hA.isHermitian.add hB.isHermitian, fun x => by rw [add_mulVec, dotProduct_add] @@ -406,21 +406,21 @@ theorem _root_.Matrix.posDef_intCast_iff [StarOrderedRing R] [DecidableEq n] [No PosDef (d : Matrix n n R) ↔ 0 < d := posDef_diagonal_iff.trans <| by simp -protected lemma add_posSemidef [CovariantClass R R (· + ·) (· ≤ · )] +protected lemma add_posSemidef [AddLeftMono R] {A : Matrix m m R} {B : Matrix m m R} (hA : A.PosDef) (hB : B.PosSemidef) : (A + B).PosDef := ⟨hA.isHermitian.add hB.isHermitian, fun x hx => by rw [add_mulVec, dotProduct_add] exact add_pos_of_pos_of_nonneg (hA.2 x hx) (hB.2 x)⟩ -protected lemma posSemidef_add [CovariantClass R R (· + ·) (· ≤ · )] +protected lemma posSemidef_add [AddLeftMono R] {A : Matrix m m R} {B : Matrix m m R} (hA : A.PosSemidef) (hB : B.PosDef) : (A + B).PosDef := ⟨hA.isHermitian.add hB.isHermitian, fun x hx => by rw [add_mulVec, dotProduct_add] exact add_pos_of_nonneg_of_pos (hA.2 x) (hB.2 x hx)⟩ -protected lemma add [CovariantClass R R (· + ·) (· ≤ · )] {A : Matrix m m R} {B : Matrix m m R} +protected lemma add [AddLeftMono R] {A : Matrix m m R} {B : Matrix m m R} (hA : A.PosDef) (hB : B.PosDef) : (A + B).PosDef := hA.add_posSemidef hB.posSemidef diff --git a/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean b/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean index 41254cc0f29bb..0a53cfbd1d352 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean @@ -1079,7 +1079,7 @@ theorem posDef_of_nonneg {Q : QuadraticMap R₂ M N} (h : ∀ x, 0 ≤ Q x) (h0 theorem posDef_iff_nonneg {Q : QuadraticMap R₂ M N} : PosDef Q ↔ (∀ x, 0 ≤ Q x) ∧ Q.Anisotropic := ⟨fun h => ⟨h.nonneg, h.anisotropic⟩, fun ⟨n, a⟩ => posDef_of_nonneg n a⟩ -theorem PosDef.add [CovariantClass N N (· + ·) (· < ·)] +theorem PosDef.add [AddLeftStrictMono N] (Q Q' : QuadraticMap R₂ M N) (hQ : PosDef Q) (hQ' : PosDef Q') : PosDef (Q + Q') := fun x hx => add_pos (hQ x hx) (hQ' x hx) diff --git a/Mathlib/LinearAlgebra/QuadraticForm/Prod.lean b/Mathlib/LinearAlgebra/QuadraticForm/Prod.lean index 20c4468b0c051..e6c65901442b7 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/Prod.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/Prod.lean @@ -144,7 +144,7 @@ theorem anisotropic_of_prod refine (h 0 x ?_).2 rw [hx, add_zero, map_zero] -theorem nonneg_prod_iff [Preorder P] [CovariantClass P P (· + ·) (· ≤ ·)] +theorem nonneg_prod_iff [Preorder P] [AddLeftMono P] {Q₁ : QuadraticMap R M₁ P} {Q₂ : QuadraticMap R M₂ P} : (∀ x, 0 ≤ (Q₁.prod Q₂) x) ↔ (∀ x, 0 ≤ Q₁ x) ∧ ∀ x, 0 ≤ Q₂ x := by simp_rw [Prod.forall, prod_apply] @@ -156,7 +156,7 @@ theorem nonneg_prod_iff [Preorder P] [CovariantClass P P (· + ·) (· ≤ ·)] · rintro ⟨h₁, h₂⟩ x₁ x₂ exact add_nonneg (h₁ x₁) (h₂ x₂) -theorem posDef_prod_iff [PartialOrder P] [CovariantClass P P (· + ·) (· ≤ ·)] +theorem posDef_prod_iff [PartialOrder P] [AddLeftMono P] {Q₁ : QuadraticMap R M₁ P} {Q₂ : QuadraticMap R M₂ P} : (Q₁.prod Q₂).PosDef ↔ Q₁.PosDef ∧ Q₂.PosDef := by simp_rw [posDef_iff_nonneg, nonneg_prod_iff] @@ -170,7 +170,7 @@ theorem posDef_prod_iff [PartialOrder P] [CovariantClass P P (· + ·) (· ≤ rw [add_eq_zero_iff_of_nonneg (hle₁ x₁) (hle₂ x₂), ha₁.eq_zero_iff, ha₂.eq_zero_iff] at hx rwa [Prod.mk_eq_zero] -theorem PosDef.prod [PartialOrder P] [CovariantClass P P (· + ·) (· ≤ ·)] +theorem PosDef.prod [PartialOrder P] [AddLeftMono P] {Q₁ : QuadraticMap R M₁ P} {Q₂ : QuadraticMap R M₂ P} (h₁ : Q₁.PosDef) (h₂ : Q₂.PosDef) : (Q₁.prod Q₂).PosDef := posDef_prod_iff.mpr ⟨h₁, h₂⟩ diff --git a/Mathlib/MeasureTheory/Function/LpOrder.lean b/Mathlib/MeasureTheory/Function/LpOrder.lean index 5a7e13be9a62b..07bfca33ec50e 100644 --- a/Mathlib/MeasureTheory/Function/LpOrder.lean +++ b/Mathlib/MeasureTheory/Function/LpOrder.lean @@ -45,7 +45,7 @@ theorem coeFn_nonneg (f : Lp E p μ) : 0 ≤ᵐ[μ] f ↔ 0 ≤ f := by · rwa [h2] · rwa [← h2] -instance instCovariantClassLE : CovariantClass (Lp E p μ) (Lp E p μ) (· + ·) (· ≤ ·) := by +instance instAddLeftMono : AddLeftMono (Lp E p μ) := by refine ⟨fun f g₁ g₂ hg₁₂ => ?_⟩ rw [← coeFn_le] at hg₁₂ ⊢ filter_upwards [coeFn_add f g₁, coeFn_add f g₂, hg₁₂] with _ h1 h2 h3 diff --git a/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean b/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean index e017f36bbf06c..920ebf5d42e0d 100644 --- a/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean +++ b/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean @@ -742,8 +742,7 @@ variable {G : Type*} [NormedLatticeAddCommGroup G] theorem coeFn_le (f g : Lp.simpleFunc G p μ) : (f : α → G) ≤ᵐ[μ] g ↔ f ≤ g := by rw [← Subtype.coe_le_coe, ← Lp.coeFn_le] -instance instCovariantClassLE : - CovariantClass (Lp.simpleFunc G p μ) (Lp.simpleFunc G p μ) (· + ·) (· ≤ ·) := by +instance instAddLeftMono : AddLeftMono (Lp.simpleFunc G p μ) := by refine ⟨fun f g₁ g₂ hg₁₂ => ?_⟩ rw [← Lp.simpleFunc.coeFn_le] at hg₁₂ ⊢ have h_add_1 : ((f + g₁ : Lp.simpleFunc G p μ) : α → G) =ᵐ[μ] (f : α → G) + g₁ := Lp.coeFn_add _ _ diff --git a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean index 3c03fb33601d8..cb05efae203e2 100644 --- a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean +++ b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean @@ -924,8 +924,7 @@ theorem lt_iff : μ < ν ↔ μ ≤ ν ∧ ∃ s, MeasurableSet s ∧ μ s < ν theorem lt_iff' : μ < ν ↔ μ ≤ ν ∧ ∃ s, μ s < ν s := lt_iff_le_not_le.trans <| and_congr Iff.rfl <| by simp only [le_iff', not_forall, not_le] -instance covariantAddLE {_ : MeasurableSpace α} : - CovariantClass (Measure α) (Measure α) (· + ·) (· ≤ ·) := +instance instAddLeftMono {_ : MeasurableSpace α} : AddLeftMono (Measure α) := ⟨fun _ν _μ₁ _μ₂ hμ s => add_le_add_left (hμ s) _⟩ protected theorem le_add_left (h : μ ≤ ν) : μ ≤ ν' + ν := fun s => le_add_left (h s) diff --git a/Mathlib/MeasureTheory/Measure/Restrict.lean b/Mathlib/MeasureTheory/Measure/Restrict.lean index 70efb6ae3f55d..128b3126b22c9 100644 --- a/Mathlib/MeasureTheory/Measure/Restrict.lean +++ b/Mathlib/MeasureTheory/Measure/Restrict.lean @@ -634,8 +634,7 @@ theorem div_ae_eq_one {β} [Group β] (f g : α → β) : f / g =ᵐ[μ] 1 ↔ f · rwa [Pi.div_apply, Pi.one_apply, div_eq_one] @[to_additive sub_nonneg_ae] -lemma one_le_div_ae {β : Type*} [Group β] [LE β] - [CovariantClass β β (Function.swap (· * ·)) (· ≤ ·)] (f g : α → β) : +lemma one_le_div_ae {β : Type*} [Group β] [LE β] [MulRightMono β] (f g : α → β) : 1 ≤ᵐ[μ] g / f ↔ f ≤ᵐ[μ] g := by refine ⟨fun h ↦ h.mono fun a ha ↦ ?_, fun h ↦ h.mono fun a ha ↦ ?_⟩ · rwa [Pi.one_apply, Pi.div_apply, one_le_div'] at ha diff --git a/Mathlib/MeasureTheory/Measure/VectorMeasure.lean b/Mathlib/MeasureTheory/Measure/VectorMeasure.lean index 06d029fc3a63b..e5e7a7d55a5d6 100644 --- a/Mathlib/MeasureTheory/Measure/VectorMeasure.lean +++ b/Mathlib/MeasureTheory/Measure/VectorMeasure.lean @@ -880,10 +880,9 @@ end section variable {M : Type*} [TopologicalSpace M] [AddCommMonoid M] [PartialOrder M] - [CovariantClass M M (· + ·) (· ≤ ·)] [ContinuousAdd M] + [AddLeftMono M] [ContinuousAdd M] -instance covariant_add_le : - CovariantClass (VectorMeasure α M) (VectorMeasure α M) (· + ·) (· ≤ ·) := +instance instAddLeftMono : AddLeftMono (VectorMeasure α M) := ⟨fun _ _ _ h i hi => add_le_add_left (h i hi) _⟩ end diff --git a/Mathlib/Order/ConditionallyCompleteLattice/Group.lean b/Mathlib/Order/ConditionallyCompleteLattice/Group.lean index a38d8ea9c6490..6d0082e85feca 100644 --- a/Mathlib/Order/ConditionallyCompleteLattice/Group.lean +++ b/Mathlib/Order/ConditionallyCompleteLattice/Group.lean @@ -19,34 +19,32 @@ variable {α : Type*} {ι : Sort*} {ι' : Sort*} [Nonempty ι] [Nonempty ι'] [ConditionallyCompleteLattice α] [Group α] @[to_additive] -theorem le_mul_ciInf [CovariantClass α α (· * ·) (· ≤ ·)] {a : α} {g : α} {h : ι → α} +theorem le_mul_ciInf [MulLeftMono α] {a : α} {g : α} {h : ι → α} (H : ∀ j, a ≤ g * h j) : a ≤ g * iInf h := inv_mul_le_iff_le_mul.mp <| le_ciInf fun _ => inv_mul_le_iff_le_mul.mpr <| H _ @[to_additive] -theorem mul_ciSup_le [CovariantClass α α (· * ·) (· ≤ ·)] {a : α} {g : α} {h : ι → α} +theorem mul_ciSup_le [MulLeftMono α] {a : α} {g : α} {h : ι → α} (H : ∀ j, g * h j ≤ a) : g * iSup h ≤ a := le_mul_ciInf (α := αᵒᵈ) H @[to_additive] -theorem le_ciInf_mul [CovariantClass α α (Function.swap (· * ·)) (· ≤ ·)] {a : α} {g : ι → α} +theorem le_ciInf_mul [MulRightMono α] {a : α} {g : ι → α} {h : α} (H : ∀ i, a ≤ g i * h) : a ≤ iInf g * h := mul_inv_le_iff_le_mul.mp <| le_ciInf fun _ => mul_inv_le_iff_le_mul.mpr <| H _ @[to_additive] -theorem ciSup_mul_le [CovariantClass α α (Function.swap (· * ·)) (· ≤ ·)] {a : α} {g : ι → α} +theorem ciSup_mul_le [MulRightMono α] {a : α} {g : ι → α} {h : α} (H : ∀ i, g i * h ≤ a) : iSup g * h ≤ a := le_ciInf_mul (α := αᵒᵈ) H @[to_additive] -theorem le_ciInf_mul_ciInf [CovariantClass α α (· * ·) (· ≤ ·)] - [CovariantClass α α (Function.swap (· * ·)) (· ≤ ·)] {a : α} {g : ι → α} {h : ι' → α} +theorem le_ciInf_mul_ciInf [MulLeftMono α] [MulRightMono α] {a : α} {g : ι → α} {h : ι' → α} (H : ∀ i j, a ≤ g i * h j) : a ≤ iInf g * iInf h := le_ciInf_mul fun _ => le_mul_ciInf <| H _ @[to_additive] -theorem ciSup_mul_ciSup_le [CovariantClass α α (· * ·) (· ≤ ·)] - [CovariantClass α α (Function.swap (· * ·)) (· ≤ ·)] {a : α} {g : ι → α} {h : ι' → α} +theorem ciSup_mul_ciSup_le [MulLeftMono α] [MulRightMono α] {a : α} {g : ι → α} {h : ι' → α} (H : ∀ i j, g i * h j ≤ a) : iSup g * iSup h ≤ a := ciSup_mul_le fun _ => mul_ciSup_le <| H _ diff --git a/Mathlib/Order/Filter/Pointwise.lean b/Mathlib/Order/Filter/Pointwise.lean index 478f544d72b32..987d9fae9f218 100644 --- a/Mathlib/Order/Filter/Pointwise.lean +++ b/Mathlib/Order/Filter/Pointwise.lean @@ -312,11 +312,11 @@ theorem le_mul_iff : h ≤ f * g ↔ ∀ ⦃s⦄, s ∈ f → ∀ ⦃t⦄, t ∈ le_map₂_iff @[to_additive] -instance covariant_mul : CovariantClass (Filter α) (Filter α) (· * ·) (· ≤ ·) := +instance mulLeftMono : MulLeftMono (Filter α) := ⟨fun _ _ _ => map₂_mono_left⟩ @[to_additive] -instance covariant_swap_mul : CovariantClass (Filter α) (Filter α) (swap (· * ·)) (· ≤ ·) := +instance mulRightMono : MulRightMono (Filter α) := ⟨fun _ _ _ => map₂_mono_right⟩ @[to_additive] diff --git a/Mathlib/Order/Filter/Ring.lean b/Mathlib/Order/Filter/Ring.lean index 3a1e434b51dd2..e6b7875f4e4ad 100644 --- a/Mathlib/Order/Filter/Ring.lean +++ b/Mathlib/Order/Filter/Ring.lean @@ -23,8 +23,8 @@ theorem EventuallyLE.mul_le_mul [MulZeroClass β] [PartialOrder β] [PosMulMono filter_upwards [hf, hg, hg₀, hf₀] with x using _root_.mul_le_mul @[to_additive EventuallyLE.add_le_add] -theorem EventuallyLE.mul_le_mul' [Mul β] [Preorder β] [CovariantClass β β (· * ·) (· ≤ ·)] - [CovariantClass β β (swap (· * ·)) (· ≤ ·)] {l : Filter α} {f₁ f₂ g₁ g₂ : α → β} +theorem EventuallyLE.mul_le_mul' [Mul β] [Preorder β] [MulLeftMono β] + [MulRightMono β] {l : Filter α} {f₁ f₂ g₁ g₂ : α → β} (hf : f₁ ≤ᶠ[l] f₂) (hg : g₁ ≤ᶠ[l] g₂) : f₁ * g₁ ≤ᶠ[l] f₂ * g₂ := by filter_upwards [hf, hg] with x hfx hgx using _root_.mul_le_mul' hfx hgx diff --git a/Mathlib/Order/LiminfLimsup.lean b/Mathlib/Order/LiminfLimsup.lean index 300e1d0160937..a83f19741bce1 100644 --- a/Mathlib/Order/LiminfLimsup.lean +++ b/Mathlib/Order/LiminfLimsup.lean @@ -305,8 +305,7 @@ lemma isBoundedUnder_sum {κ : Type*} [AddCommMonoid R] {r : R → R → Prop} variable [Preorder R] -lemma isBoundedUnder_ge_add [Add R] - [CovariantClass R R (fun a b ↦ a + b) (· ≤ ·)] [CovariantClass R R (fun a b ↦ b + a) (· ≤ ·)] +lemma isBoundedUnder_ge_add [Add R] [AddLeftMono R] [AddRightMono R] {u v : α → R} (u_bdd_ge : f.IsBoundedUnder (· ≥ ·) u) (v_bdd_ge : f.IsBoundedUnder (· ≥ ·) v) : f.IsBoundedUnder (· ≥ ·) (u + v) := by obtain ⟨U, hU⟩ := u_bdd_ge @@ -315,8 +314,7 @@ lemma isBoundedUnder_ge_add [Add R] simp only [eventually_map, Pi.add_apply] at hU hV ⊢ filter_upwards [hU, hV] with a hu hv using add_le_add hu hv -lemma isBoundedUnder_le_add [Add R] - [CovariantClass R R (fun a b ↦ a + b) (· ≤ ·)] [CovariantClass R R (fun a b ↦ b + a) (· ≤ ·)] +lemma isBoundedUnder_le_add [Add R] [AddLeftMono R] [AddRightMono R] {u v : α → R} (u_bdd_le : f.IsBoundedUnder (· ≤ ·) u) (v_bdd_le : f.IsBoundedUnder (· ≤ ·) v) : f.IsBoundedUnder (· ≤ ·) (u + v) := by obtain ⟨U, hU⟩ := u_bdd_le @@ -325,14 +323,12 @@ lemma isBoundedUnder_le_add [Add R] simp only [eventually_map, Pi.add_apply] at hU hV ⊢ filter_upwards [hU, hV] with a hu hv using add_le_add hu hv -lemma isBoundedUnder_le_sum {κ : Type*} [AddCommMonoid R] - [CovariantClass R R (fun a b ↦ a + b) (· ≤ ·)] [CovariantClass R R (fun a b ↦ b + a) (· ≤ ·)] +lemma isBoundedUnder_le_sum {κ : Type*} [AddCommMonoid R] [AddLeftMono R] [AddRightMono R] {u : κ → α → R} (s : Finset κ) : (∀ k ∈ s, f.IsBoundedUnder (· ≤ ·) (u k)) → f.IsBoundedUnder (· ≤ ·) (∑ k ∈ s, u k) := by apply isBoundedUnder_sum (fun _ _ ↦ isBoundedUnder_le_add) le_rfl -lemma isBoundedUnder_ge_sum {κ : Type*} [AddCommMonoid R] - [CovariantClass R R (fun a b ↦ a + b) (· ≤ ·)] [CovariantClass R R (fun a b ↦ b + a) (· ≤ ·)] +lemma isBoundedUnder_ge_sum {κ : Type*} [AddCommMonoid R] [AddLeftMono R] [AddRightMono R] {u : κ → α → R} (s : Finset κ) : (∀ k ∈ s, f.IsBoundedUnder (· ≥ ·) (u k)) → f.IsBoundedUnder (· ≥ ·) (∑ k ∈ s, u k) := by diff --git a/Mathlib/Probability/Martingale/Basic.lean b/Mathlib/Probability/Martingale/Basic.lean index 08032df92e947..1693e5117c3b3 100644 --- a/Mathlib/Probability/Martingale/Basic.lean +++ b/Mathlib/Probability/Martingale/Basic.lean @@ -161,7 +161,7 @@ theorem setIntegral_le [SigmaFiniteFiltration μ ℱ] {f : ι → Ω → ℝ} (h @[deprecated (since := "2024-04-17")] alias set_integral_le := setIntegral_le -theorem add [Preorder E] [CovariantClass E E (· + ·) (· ≤ ·)] (hf : Supermartingale f ℱ μ) +theorem add [Preorder E] [AddLeftMono E] (hf : Supermartingale f ℱ μ) (hg : Supermartingale g ℱ μ) : Supermartingale (f + g) ℱ μ := by refine ⟨hf.1.add hg.1, fun i j hij => ?_, fun i => (hf.2.2 i).add (hg.2.2 i)⟩ refine (condexp_add (hf.integrable j) (hg.integrable j)).le.trans ?_ @@ -169,11 +169,11 @@ theorem add [Preorder E] [CovariantClass E E (· + ·) (· ≤ ·)] (hf : Superm intros refine add_le_add ?_ ?_ <;> assumption -theorem add_martingale [Preorder E] [CovariantClass E E (· + ·) (· ≤ ·)] +theorem add_martingale [Preorder E] [AddLeftMono E] (hf : Supermartingale f ℱ μ) (hg : Martingale g ℱ μ) : Supermartingale (f + g) ℱ μ := hf.add hg.supermartingale -theorem neg [Preorder E] [CovariantClass E E (· + ·) (· ≤ ·)] (hf : Supermartingale f ℱ μ) : +theorem neg [Preorder E] [AddLeftMono E] (hf : Supermartingale f ℱ μ) : Submartingale (-f) ℱ μ := by refine ⟨hf.1.neg, fun i j hij => ?_, fun i => (hf.2.2 i).neg⟩ refine EventuallyLE.trans ?_ (condexp_neg (f j)).symm.le @@ -198,7 +198,7 @@ theorem ae_le_condexp [LE E] (hf : Submartingale f ℱ μ) {i j : ι} (hij : i f i ≤ᵐ[μ] μ[f j|ℱ i] := hf.2.1 i j hij -theorem add [Preorder E] [CovariantClass E E (· + ·) (· ≤ ·)] (hf : Submartingale f ℱ μ) +theorem add [Preorder E] [AddLeftMono E] (hf : Submartingale f ℱ μ) (hg : Submartingale g ℱ μ) : Submartingale (f + g) ℱ μ := by refine ⟨hf.1.add hg.1, fun i j hij => ?_, fun i => (hf.2.2 i).add (hg.2.2 i)⟩ refine EventuallyLE.trans ?_ (condexp_add (hf.integrable j) (hg.integrable j)).symm.le @@ -206,11 +206,11 @@ theorem add [Preorder E] [CovariantClass E E (· + ·) (· ≤ ·)] (hf : Submar intros refine add_le_add ?_ ?_ <;> assumption -theorem add_martingale [Preorder E] [CovariantClass E E (· + ·) (· ≤ ·)] (hf : Submartingale f ℱ μ) +theorem add_martingale [Preorder E] [AddLeftMono E] (hf : Submartingale f ℱ μ) (hg : Martingale g ℱ μ) : Submartingale (f + g) ℱ μ := hf.add hg.submartingale -theorem neg [Preorder E] [CovariantClass E E (· + ·) (· ≤ ·)] (hf : Submartingale f ℱ μ) : +theorem neg [Preorder E] [AddLeftMono E] (hf : Submartingale f ℱ μ) : Supermartingale (-f) ℱ μ := by refine ⟨hf.1.neg, fun i j hij => (condexp_neg (f j)).le.trans ?_, fun i => (hf.2.2 i).neg⟩ filter_upwards [hf.2.1 i j hij] with _ _ @@ -226,11 +226,11 @@ theorem setIntegral_le [SigmaFiniteFiltration μ ℱ] {f : ι → Ω → ℝ} (h @[deprecated (since := "2024-04-17")] alias set_integral_le := setIntegral_le -theorem sub_supermartingale [Preorder E] [CovariantClass E E (· + ·) (· ≤ ·)] +theorem sub_supermartingale [Preorder E] [AddLeftMono E] (hf : Submartingale f ℱ μ) (hg : Supermartingale g ℱ μ) : Submartingale (f - g) ℱ μ := by rw [sub_eq_add_neg]; exact hf.add hg.neg -theorem sub_martingale [Preorder E] [CovariantClass E E (· + ·) (· ≤ ·)] (hf : Submartingale f ℱ μ) +theorem sub_martingale [Preorder E] [AddLeftMono E] (hf : Submartingale f ℱ μ) (hg : Martingale g ℱ μ) : Submartingale (f - g) ℱ μ := hf.sub_supermartingale hg.supermartingale @@ -299,11 +299,11 @@ end Submartingale namespace Supermartingale -theorem sub_submartingale [Preorder E] [CovariantClass E E (· + ·) (· ≤ ·)] +theorem sub_submartingale [Preorder E] [AddLeftMono E] (hf : Supermartingale f ℱ μ) (hg : Submartingale g ℱ μ) : Supermartingale (f - g) ℱ μ := by rw [sub_eq_add_neg]; exact hf.add hg.neg -theorem sub_martingale [Preorder E] [CovariantClass E E (· + ·) (· ≤ ·)] +theorem sub_martingale [Preorder E] [AddLeftMono E] (hf : Supermartingale f ℱ μ) (hg : Martingale g ℱ μ) : Supermartingale (f - g) ℱ μ := hf.sub_submartingale hg.submartingale diff --git a/Mathlib/Probability/Process/Stopping.lean b/Mathlib/Probability/Process/Stopping.lean index a68d221da2416..c9b88bfbb13b3 100644 --- a/Mathlib/Probability/Process/Stopping.lean +++ b/Mathlib/Probability/Process/Stopping.lean @@ -245,8 +245,8 @@ protected theorem min_const [LinearOrder ι] {f : Filtration ι m} {τ : Ω → (hτ : IsStoppingTime f τ) (i : ι) : IsStoppingTime f fun ω => min (τ ω) i := hτ.min (isStoppingTime_const f i) -theorem add_const [AddGroup ι] [Preorder ι] [CovariantClass ι ι (Function.swap (· + ·)) (· ≤ ·)] - [CovariantClass ι ι (· + ·) (· ≤ ·)] {f : Filtration ι m} {τ : Ω → ι} (hτ : IsStoppingTime f τ) +theorem add_const [AddGroup ι] [Preorder ι] [AddRightMono ι] + [AddLeftMono ι] {f : Filtration ι m} {τ : Ω → ι} (hτ : IsStoppingTime f τ) {i : ι} (hi : 0 ≤ i) : IsStoppingTime f fun ω => τ ω + i := by intro j simp_rw [← le_sub_iff_add_le] diff --git a/Mathlib/RingTheory/GradedAlgebra/Basic.lean b/Mathlib/RingTheory/GradedAlgebra/Basic.lean index 0d6693692b34f..0121091f00ac2 100644 --- a/Mathlib/RingTheory/GradedAlgebra/Basic.lean +++ b/Mathlib/RingTheory/GradedAlgebra/Basic.lean @@ -310,7 +310,7 @@ theorem coe_decompose_mul_of_right_mem_of_not_le (b_mem : b ∈ 𝒜 i) (h : ¬i lift b to 𝒜 i using b_mem rwa [decompose_mul, decompose_coe, coe_mul_of_apply_of_not_le] -variable [Sub ι] [OrderedSub ι] [ContravariantClass ι ι (· + ·) (· ≤ ·)] +variable [Sub ι] [OrderedSub ι] [AddLeftReflectLE ι] theorem coe_decompose_mul_of_left_mem_of_le (a_mem : a ∈ 𝒜 i) (h : i ≤ n) : (decompose 𝒜 (a * b) n : A) = a * decompose 𝒜 b (n - i) := by diff --git a/Mathlib/SetTheory/Cardinal/Basic.lean b/Mathlib/SetTheory/Cardinal/Basic.lean index c5202a59be16d..e44825db47401 100644 --- a/Mathlib/SetTheory/Cardinal/Basic.lean +++ b/Mathlib/SetTheory/Cardinal/Basic.lean @@ -610,10 +610,10 @@ protected theorem zero_le : ∀ a : Cardinal, 0 ≤ a := by private theorem add_le_add' : ∀ {a b c d : Cardinal}, a ≤ b → c ≤ d → a + c ≤ b + d := by rintro ⟨α⟩ ⟨β⟩ ⟨γ⟩ ⟨δ⟩ ⟨e₁⟩ ⟨e₂⟩; exact ⟨e₁.sumMap e₂⟩ -instance add_covariantClass : CovariantClass Cardinal Cardinal (· + ·) (· ≤ ·) := +instance addLeftMono : AddLeftMono Cardinal := ⟨fun _ _ _ => add_le_add' le_rfl⟩ -instance add_swap_covariantClass : CovariantClass Cardinal Cardinal (swap (· + ·)) (· ≤ ·) := +instance addRightMono : AddRightMono Cardinal := ⟨fun _ _ _ h => add_le_add' h le_rfl⟩ instance canonicallyOrderedCommSemiring : CanonicallyOrderedCommSemiring Cardinal.{u} := diff --git a/Mathlib/SetTheory/Game/Basic.lean b/Mathlib/SetTheory/Game/Basic.lean index 3504ce49ac81a..c7ac5c82b8191 100644 --- a/Mathlib/SetTheory/Game/Basic.lean +++ b/Mathlib/SetTheory/Game/Basic.lean @@ -164,22 +164,22 @@ namespace Game local infixl:50 " ⧏ " => LF local infixl:50 " ‖ " => Fuzzy -instance covariantClass_add_le : CovariantClass Game Game (· + ·) (· ≤ ·) := +instance addLeftMono : AddLeftMono Game := ⟨by rintro ⟨a⟩ ⟨b⟩ ⟨c⟩ h exact @add_le_add_left _ _ _ _ b c h a⟩ -instance covariantClass_swap_add_le : CovariantClass Game Game (swap (· + ·)) (· ≤ ·) := +instance addRightMono : AddRightMono Game := ⟨by rintro ⟨a⟩ ⟨b⟩ ⟨c⟩ h exact @add_le_add_right _ _ _ _ b c h a⟩ -instance covariantClass_add_lt : CovariantClass Game Game (· + ·) (· < ·) := +instance addLeftStrictMono : AddLeftStrictMono Game := ⟨by rintro ⟨a⟩ ⟨b⟩ ⟨c⟩ h exact @add_lt_add_left _ _ _ _ b c h a⟩ -instance covariantClass_swap_add_lt : CovariantClass Game Game (swap (· + ·)) (· < ·) := +instance addRightStrictMono : AddRightStrictMono Game := ⟨by rintro ⟨a⟩ ⟨b⟩ ⟨c⟩ h exact @add_lt_add_right _ _ _ _ b c h a⟩ @@ -194,7 +194,7 @@ theorem add_lf_add_left : ∀ {b c : Game} (_ : b ⧏ c) (a), (a + b : Game) ⧏ instance orderedAddCommGroup : OrderedAddCommGroup Game := { Game.instAddCommGroupWithOneGame, Game.instPartialOrderGame with - add_le_add_left := @add_le_add_left _ _ _ Game.covariantClass_add_le } + add_le_add_left := @add_le_add_left _ _ _ Game.addLeftMono } /-- A small family of games is bounded above. -/ lemma bddAbove_range_of_small {ι : Type*} [Small.{u} ι] (f : ι → Game.{u}) : diff --git a/Mathlib/SetTheory/Game/PGame.lean b/Mathlib/SetTheory/Game/PGame.lean index 56ee1377443a3..f67af28a88468 100644 --- a/Mathlib/SetTheory/Game/PGame.lean +++ b/Mathlib/SetTheory/Game/PGame.lean @@ -1547,10 +1547,10 @@ private theorem add_le_add_right' : ∀ {x y z : PGame}, x ≤ y → x + z ≤ y Or.inr ⟨@toRightMovesAdd _ ⟨_, _, _, _⟩ (Sum.inr i), add_le_add_right' h⟩ termination_by x y z => (x, y, z) -instance covariantClass_swap_add_le : CovariantClass PGame PGame (swap (· + ·)) (· ≤ ·) := +instance addRightMono : AddRightMono PGame := ⟨fun _ _ _ => add_le_add_right'⟩ -instance covariantClass_add_le : CovariantClass PGame PGame (· + ·) (· ≤ ·) := +instance addLeftMono : AddLeftMono PGame := ⟨fun x _ _ h => (add_comm_le.trans (add_le_add_right h x)).trans add_comm_le⟩ theorem add_lf_add_right {y z : PGame} (h : y ⧏ z) (x) : y + x ⧏ z + x := @@ -1571,10 +1571,10 @@ theorem add_lf_add_left {y z : PGame} (h : y ⧏ z) (x) : x + y ⧏ x + z := by rw [lf_congr add_comm_equiv add_comm_equiv] apply add_lf_add_right h -instance covariantClass_swap_add_lt : CovariantClass PGame PGame (swap (· + ·)) (· < ·) := +instance addRightStrictMono : AddRightStrictMono PGame := ⟨fun x _ _ h => ⟨add_le_add_right h.1 x, add_lf_add_right h.2 x⟩⟩ -instance covariantClass_add_lt : CovariantClass PGame PGame (· + ·) (· < ·) := +instance addLeftStrictMono : AddLeftStrictMono PGame := ⟨fun x _ _ h => ⟨add_le_add_left h.1 x, add_lf_add_left h.2 x⟩⟩ theorem add_lf_add_of_lf_of_le {w x y z : PGame} (hwx : w ⧏ x) (hyz : y ≤ z) : w + y ⧏ x + z := diff --git a/Mathlib/SetTheory/Ordinal/Arithmetic.lean b/Mathlib/SetTheory/Ordinal/Arithmetic.lean index 15e30c17e0a6f..6e9b1070b8f6c 100644 --- a/Mathlib/SetTheory/Ordinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Ordinal/Arithmetic.lean @@ -77,8 +77,8 @@ theorem lift_succ (a : Ordinal.{v}) : lift.{u} (succ a) = succ (lift.{u} a) := b rw [← add_one_eq_succ, lift_add, lift_one] rfl -instance add_contravariantClass_le : - ContravariantClass Ordinal.{u} Ordinal.{u} (· + ·) (· ≤ ·) where +instance instAddLeftReflectLE : + AddLeftReflectLE Ordinal.{u} where elim c a b := by refine inductionOn₃ a b c fun α r _ β s _ γ t _ ⟨f⟩ ↦ ?_ have H₁ a : f (Sum.inl a) = Sum.inl a := by @@ -99,14 +99,13 @@ theorem add_left_cancel (a) {b c : Ordinal} : a + b = a + c ↔ b = c := by private theorem add_lt_add_iff_left' (a) {b c : Ordinal} : a + b < a + c ↔ b < c := by rw [← not_le, ← not_le, add_le_add_iff_left] -instance add_covariantClass_lt : CovariantClass Ordinal.{u} Ordinal.{u} (· + ·) (· < ·) := +instance instAddLeftStrictMono : AddLeftStrictMono Ordinal.{u} := ⟨fun a _b _c ↦ (add_lt_add_iff_left' a).2⟩ -instance add_contravariantClass_lt : ContravariantClass Ordinal.{u} Ordinal.{u} (· + ·) (· < ·) := +instance instAddLeftReflectLT : AddLeftReflectLT Ordinal.{u} := ⟨fun a _b _c ↦ (add_lt_add_iff_left' a).1⟩ -instance add_swap_contravariantClass_lt : - ContravariantClass Ordinal.{u} Ordinal.{u} (swap (· + ·)) (· < ·) := +instance instAddRightReflectLT : AddRightReflectLT Ordinal.{u} := ⟨fun _a _b _c ↦ lt_imp_lt_of_le_imp_le fun h => add_le_add_right h _⟩ theorem add_le_add_iff_right {a b : Ordinal} : ∀ n : ℕ, a + n ≤ b + n ↔ a ≤ b @@ -682,7 +681,7 @@ instance leftDistribClass : LeftDistribClass Ordinal.{u} := theorem mul_succ (a b : Ordinal) : a * succ b = a * b + a := mul_add_one a b -instance mul_covariantClass_le : CovariantClass Ordinal.{u} Ordinal.{u} (· * ·) (· ≤ ·) := +instance mulLeftMono : MulLeftMono Ordinal.{u} := ⟨fun c a b => Quotient.inductionOn₃ a b c fun ⟨α, r, _⟩ ⟨β, s, _⟩ ⟨γ, t, _⟩ ⟨f⟩ => by refine @@ -691,8 +690,7 @@ instance mul_covariantClass_le : CovariantClass Ordinal.{u} Ordinal.{u} (· * · · exact Prod.Lex.left _ _ (f.toRelEmbedding.map_rel_iff.2 h') · exact Prod.Lex.right _ h'⟩ -instance mul_swap_covariantClass_le : - CovariantClass Ordinal.{u} Ordinal.{u} (swap (· * ·)) (· ≤ ·) := +instance mulRightMono : MulRightMono Ordinal.{u} := ⟨fun c a b => Quotient.inductionOn₃ a b c fun ⟨α, r, _⟩ ⟨β, s, _⟩ ⟨γ, t, _⟩ ⟨f⟩ => by refine diff --git a/Mathlib/SetTheory/Ordinal/Basic.lean b/Mathlib/SetTheory/Ordinal/Basic.lean index f3bb4990f67f0..abdf8207a1c0c 100644 --- a/Mathlib/SetTheory/Ordinal/Basic.lean +++ b/Mathlib/SetTheory/Ordinal/Basic.lean @@ -789,14 +789,13 @@ theorem card_ofNat (n : ℕ) [n.AtLeastTwo] : card.{u} (no_index (OfNat.ofNat n)) = OfNat.ofNat n := card_nat n -instance add_covariantClass_le : CovariantClass Ordinal.{u} Ordinal.{u} (· + ·) (· ≤ ·) where +instance instAddLeftMono : AddLeftMono Ordinal.{u} where elim c a b := by refine inductionOn₃ a b c fun α r _ β s _ γ t _ ⟨f⟩ ↦ (RelEmbedding.ofMonotone (Sum.recOn · Sum.inl (Sum.inr ∘ f)) ?_).ordinal_type_le simp [f.map_rel_iff] -instance add_swap_covariantClass_le : - CovariantClass Ordinal.{u} Ordinal.{u} (swap (· + ·)) (· ≤ ·) where +instance instAddRightMono : AddRightMono Ordinal.{u} where elim c a b := by refine inductionOn₃ a b c fun α r _ β s _ γ t _ ⟨f⟩ ↦ (RelEmbedding.ofMonotone (Sum.recOn · (Sum.inl ∘ f) Sum.inr) ?_).ordinal_type_le diff --git a/Mathlib/SetTheory/Ordinal/NaturalOps.lean b/Mathlib/SetTheory/Ordinal/NaturalOps.lean index 85ff9d0c970cb..e4d93ce91b40d 100644 --- a/Mathlib/SetTheory/Ordinal/NaturalOps.lean +++ b/Mathlib/SetTheory/Ordinal/NaturalOps.lean @@ -319,14 +319,13 @@ open Ordinal NaturalOps instance : Add NatOrdinal := ⟨nadd⟩ instance : SuccAddOrder NatOrdinal := ⟨fun x => (nadd_one x).symm⟩ -instance add_covariantClass_lt : CovariantClass NatOrdinal.{u} NatOrdinal.{u} (· + ·) (· < ·) := +instance addLeftStrictMono : AddLeftStrictMono NatOrdinal.{u} := ⟨fun a _ _ h => nadd_lt_nadd_left h a⟩ -instance add_covariantClass_le : CovariantClass NatOrdinal.{u} NatOrdinal.{u} (· + ·) (· ≤ ·) := +instance addLeftMono : AddLeftMono NatOrdinal.{u} := ⟨fun a _ _ h => nadd_le_nadd_left h a⟩ -instance add_contravariantClass_le : - ContravariantClass NatOrdinal.{u} NatOrdinal.{u} (· + ·) (· ≤ ·) := +instance addLeftReflectLE : AddLeftReflectLE NatOrdinal.{u} := ⟨fun a b c h => by by_contra! h' exact h.not_lt (add_lt_add_left h' a)⟩ diff --git a/Mathlib/Tactic/Bound.lean b/Mathlib/Tactic/Bound.lean index 47591bb967e5d..58f217d6fa30f 100644 --- a/Mathlib/Tactic/Bound.lean +++ b/Mathlib/Tactic/Bound.lean @@ -112,7 +112,7 @@ lemma Nat.cast_pos_of_pos {R : Type} [OrderedSemiring R] [Nontrivial R] {n : ℕ Nat.cast_pos.mpr lemma Nat.one_le_cast_of_le {α : Type} [AddCommMonoidWithOne α] [PartialOrder α] - [CovariantClass α α (fun (x y : α) => x + y) fun (x y : α) => x ≤ y] [ZeroLEOneClass α] + [AddLeftMono α] [ZeroLEOneClass α] [CharZero α] {n : ℕ} : 1 ≤ n → 1 ≤ (n : α) := Nat.one_le_cast.mpr diff --git a/Mathlib/Tactic/GCongr/Core.lean b/Mathlib/Tactic/GCongr/Core.lean index bf7982ef2748e..820c5f95c16f2 100644 --- a/Mathlib/Tactic/GCongr/Core.lean +++ b/Mathlib/Tactic/GCongr/Core.lean @@ -78,8 +78,8 @@ functions which have different theories when different typeclass assumptions app the following lemma is stored with the same `@[gcongr]` data as `mul_le_mul` above, and the two lemmas are simply tried in succession to determine which has the typeclasses relevant to the goal: ``` -theorem mul_le_mul' [Mul α] [Preorder α] [CovariantClass α α (· * ·) (· ≤ ·)] - [CovariantClass α α (Function.swap (· * ·)) (· ≤ ·)] {a b c d : α} (h₁ : a ≤ b) (h₂ : c ≤ d) : +theorem mul_le_mul' [Mul α] [Preorder α] [MulLeftMono α] + [MulRightMono α] {a b c d : α} (h₁ : a ≤ b) (h₂ : c ≤ d) : a * c ≤ b * d ``` diff --git a/Mathlib/Tactic/Positivity/Basic.lean b/Mathlib/Tactic/Positivity/Basic.lean index d8eeaa14347fa..6d1cf31d771de 100644 --- a/Mathlib/Tactic/Positivity/Basic.lean +++ b/Mathlib/Tactic/Positivity/Basic.lean @@ -159,16 +159,16 @@ such that `positivity` successfully recognises both `a` and `b`. -/ let ra ← core zα pα a; let rb ← core zα pα b match ra, rb with | .positive pa, .positive pb => - let _a ← synthInstanceQ (q(CovariantClass $α $α (·+·) (·<·)) : Q(Prop)) + let _a ← synthInstanceQ (q(AddLeftStrictMono $α) : Q(Prop)) pure (.positive q(add_pos $pa $pb)) | .positive pa, .nonnegative pb => - let _a ← synthInstanceQ (q(CovariantClass $α $α (swap (·+·)) (·<·)) : Q(Prop)) + let _a ← synthInstanceQ (q(AddRightStrictMono $α) : Q(Prop)) pure (.positive q(lt_add_of_pos_of_le $pa $pb)) | .nonnegative pa, .positive pb => - let _a ← synthInstanceQ (q(CovariantClass $α $α (·+·) (·<·)) : Q(Prop)) + let _a ← synthInstanceQ (q(AddLeftStrictMono $α) : Q(Prop)) pure (.positive q(lt_add_of_le_of_pos $pa $pb)) | .nonnegative pa, .nonnegative pb => - let _a ← synthInstanceQ (q(CovariantClass $α $α (·+·) (·≤·)) : Q(Prop)) + let _a ← synthInstanceQ (q(AddLeftMono $α) : Q(Prop)) pure (.nonnegative q(add_nonneg $pa $pb)) | _, _ => failure @@ -304,7 +304,7 @@ def evalPow : PositivityExt where eval {u α} zα pα e := do | .none => pure .none private theorem abs_pos_of_ne_zero {α : Type*} [AddGroup α] [LinearOrder α] - [CovariantClass α α (·+·) (·≤·)] {a : α} : a ≠ 0 → 0 < |a| := abs_pos.mpr + [AddLeftMono α] {a : α} : a ≠ 0 → 0 < |a| := abs_pos.mpr /-- The `positivity` extension which identifies expressions of the form `|a|`. -/ @[positivity |_|] diff --git a/Mathlib/Tactic/Ring/Compare.lean b/Mathlib/Tactic/Ring/Compare.lean index 5868ee2ee1706..967a8599abab2 100644 --- a/Mathlib/Tactic/Ring/Compare.lean +++ b/Mathlib/Tactic/Ring/Compare.lean @@ -66,7 +66,7 @@ abbrev lt_of_socs (α : Type*) [StrictOrderedCommSemiring α] : LT α := inferIn end Typeclass /-! The lemmas like `add_le_add_right` in the root namespace are stated under minimal type classes, -typically just `[CovariantClass α α (swap (· + ·)) (· ≤ ·)]` or similar. Here we restate these +typically just `[AddRightMono α]` or similar. Here we restate these lemmas under stronger type class assumptions (`[OrderedCommSemiring α]` or similar), which helps in speeding up the metaprograms in this file (`Mathlib.Tactic.Ring.proveLT` and `Mathlib.Tactic.Ring.proveLE`) substantially -- about a 50% reduction in heartbeat count in diff --git a/Mathlib/Topology/Algebra/Order/LiminfLimsup.lean b/Mathlib/Topology/Algebra/Order/LiminfLimsup.lean index 72c9ec5be2a80..18a175654b3dd 100644 --- a/Mathlib/Topology/Algebra/Order/LiminfLimsup.lean +++ b/Mathlib/Topology/Algebra/Order/LiminfLimsup.lean @@ -525,7 +525,7 @@ variable [ConditionallyCompleteLinearOrder R] [TopologicalSpace R] [OrderTopolog /-- `liminf (c + xᵢ) = c + liminf xᵢ`. -/ lemma limsup_const_add (F : Filter ι) [NeBot F] [Add R] [ContinuousAdd R] - [CovariantClass R R (fun x y ↦ x + y) fun x y ↦ x ≤ y] (f : ι → R) (c : R) + [AddLeftMono R] (f : ι → R) (c : R) (bdd_above : F.IsBoundedUnder (· ≤ ·) f) (cobdd : F.IsCoboundedUnder (· ≤ ·) f) : Filter.limsup (fun i ↦ c + f i) F = c + Filter.limsup f F := (Monotone.map_limsSup_of_continuousAt (F := F.map f) (f := fun (x : R) ↦ c + x) @@ -533,7 +533,7 @@ lemma limsup_const_add (F : Filter ι) [NeBot F] [Add R] [ContinuousAdd R] /-- `limsup (xᵢ + c) = (limsup xᵢ) + c`. -/ lemma limsup_add_const (F : Filter ι) [NeBot F] [Add R] [ContinuousAdd R] - [CovariantClass R R (Function.swap fun x y ↦ x + y) fun x y ↦ x ≤ y] (f : ι → R) (c : R) + [AddRightMono R] (f : ι → R) (c : R) (bdd_above : F.IsBoundedUnder (· ≤ ·) f) (cobdd : F.IsCoboundedUnder (· ≤ ·) f) : Filter.limsup (fun i ↦ f i + c) F = Filter.limsup f F + c := (Monotone.map_limsSup_of_continuousAt (F := F.map f) (f := fun (x : R) ↦ x + c) @@ -541,7 +541,7 @@ lemma limsup_add_const (F : Filter ι) [NeBot F] [Add R] [ContinuousAdd R] /-- `liminf (c + xᵢ) = c + limsup xᵢ`. -/ lemma liminf_const_add (F : Filter ι) [NeBot F] [Add R] [ContinuousAdd R] - [CovariantClass R R (fun x y ↦ x + y) fun x y ↦ x ≤ y] (f : ι → R) (c : R) + [AddLeftMono R] (f : ι → R) (c : R) (cobdd : F.IsCoboundedUnder (· ≥ ·) f) (bdd_below : F.IsBoundedUnder (· ≥ ·) f) : Filter.liminf (fun i ↦ c + f i) F = c + Filter.liminf f F := (Monotone.map_limsInf_of_continuousAt (F := F.map f) (f := fun (x : R) ↦ c + x) @@ -549,7 +549,7 @@ lemma liminf_const_add (F : Filter ι) [NeBot F] [Add R] [ContinuousAdd R] /-- `liminf (xᵢ + c) = (liminf xᵢ) + c`. -/ lemma liminf_add_const (F : Filter ι) [NeBot F] [Add R] [ContinuousAdd R] - [CovariantClass R R (Function.swap fun x y ↦ x + y) fun x y ↦ x ≤ y] (f : ι → R) (c : R) + [AddRightMono R] (f : ι → R) (c : R) (cobdd : F.IsCoboundedUnder (· ≥ ·) f) (bdd_below : F.IsBoundedUnder (· ≥ ·) f) : Filter.liminf (fun i ↦ f i + c) F = Filter.liminf f F + c := (Monotone.map_limsInf_of_continuousAt (F := F.map f) (f := fun (x : R) ↦ x + c) @@ -557,7 +557,7 @@ lemma liminf_add_const (F : Filter ι) [NeBot F] [Add R] [ContinuousAdd R] /-- `limsup (c - xᵢ) = c - liminf xᵢ`. -/ lemma limsup_const_sub (F : Filter ι) [AddCommSemigroup R] [Sub R] [ContinuousSub R] [OrderedSub R] - [CovariantClass R R (fun x y ↦ x + y) fun x y ↦ x ≤ y] (f : ι → R) (c : R) + [AddLeftMono R] (f : ι → R) (c : R) (cobdd : F.IsCoboundedUnder (· ≥ ·) f) (bdd_below : F.IsBoundedUnder (· ≥ ·) f) : Filter.limsup (fun i ↦ c - f i) F = c - Filter.liminf f F := by rcases F.eq_or_neBot with rfl | _ @@ -594,7 +594,7 @@ lemma limsup_sub_const (F : Filter ι) [AddCommSemigroup R] [Sub R] [ContinuousS /-- `liminf (c - xᵢ) = c - limsup xᵢ`. -/ lemma liminf_const_sub (F : Filter ι) [NeBot F] [AddCommSemigroup R] [Sub R] [ContinuousSub R] - [OrderedSub R] [CovariantClass R R (fun x y ↦ x + y) fun x y ↦ x ≤ y] (f : ι → R) (c : R) + [OrderedSub R] [AddLeftMono R] (f : ι → R) (c : R) (bdd_above : F.IsBoundedUnder (· ≤ ·) f) (cobdd : F.IsCoboundedUnder (· ≤ ·) f) : Filter.liminf (fun i ↦ c - f i) F = c - Filter.limsup f F := (Antitone.map_limsSup_of_continuousAt (F := F.map f) (f := fun (x : R) ↦ c - x) diff --git a/Mathlib/Topology/ContinuousMap/Algebra.lean b/Mathlib/Topology/ContinuousMap/Algebra.lean index f183492f6c269..c475ae94f7dc9 100644 --- a/Mathlib/Topology/ContinuousMap/Algebra.lean +++ b/Mathlib/Topology/ContinuousMap/Algebra.lean @@ -832,15 +832,13 @@ variable {β : Type*} [TopologicalSpace β] /-! `C(α, β)`is a lattice ordered group -/ @[to_additive] -instance instCovariantClass_mul_le_left [PartialOrder β] [Mul β] [ContinuousMul β] - [CovariantClass β β (· * ·) (· ≤ ·)] : - CovariantClass C(α, β) C(α, β) (· * ·) (· ≤ ·) := +instance instMulLeftMono [PartialOrder β] [Mul β] [ContinuousMul β] [MulLeftMono β] : + MulLeftMono C(α, β) := ⟨fun _ _ _ hg₁₂ x => mul_le_mul_left' (hg₁₂ x) _⟩ @[to_additive] -instance instCovariantClass_mul_le_right [PartialOrder β] [Mul β] [ContinuousMul β] - [CovariantClass β β (Function.swap (· * ·)) (· ≤ ·)] : - CovariantClass C(α, β) C(α, β) (Function.swap (· * ·)) (· ≤ ·) := +instance instMulRightMono [PartialOrder β] [Mul β] [ContinuousMul β] [MulRightMono β] : + MulRightMono C(α, β) := ⟨fun _ _ _ hg₁₂ x => mul_le_mul_right' (hg₁₂ x) _⟩ variable [Group β] [TopologicalGroup β] [Lattice β] [TopologicalLattice β] diff --git a/Mathlib/Topology/Support.lean b/Mathlib/Topology/Support.lean index f0d049776a387..30537f87b604c 100644 --- a/Mathlib/Topology/Support.lean +++ b/Mathlib/Topology/Support.lean @@ -362,8 +362,7 @@ end MulZeroClass section OrderedAddGroup -variable [TopologicalSpace α] [AddGroup β] [Lattice β] - [CovariantClass β β (· + ·) (· ≤ ·)] +variable [TopologicalSpace α] [AddGroup β] [Lattice β] [AddLeftMono β] protected theorem HasCompactSupport.abs {f : α → β} (hf : HasCompactSupport f) : HasCompactSupport |f| := From 3b956ec74be6771f25956f90740af8ded0572b70 Mon Sep 17 00:00:00 2001 From: Daniel Carranza <61289743+daniel-carranza@users.noreply.github.com> Date: Thu, 24 Oct 2024 22:12:00 +0000 Subject: [PATCH 396/505] feat(CategoryTheory/Closed/Enrichment): a closed monoidal category is enriched in itself (#17326) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit From a monoidal closed category V, construct an instance of `EnrichedCategory V` on V where the hom-objects are given by the internal homs of the closed structure. Co-authored-by: Joël Riou <37772949+joelriou@users.noreply.github.com> Co-authored-by: daniel-carranza --- Mathlib.lean | 1 + Mathlib/CategoryTheory/Closed/Enrichment.lean | 41 +++++++++ Mathlib/CategoryTheory/Closed/Monoidal.lean | 83 +++++++++++++++++++ 3 files changed, 125 insertions(+) create mode 100644 Mathlib/CategoryTheory/Closed/Enrichment.lean diff --git a/Mathlib.lean b/Mathlib.lean index 6d9a2bf44e45d..e839e2377857a 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1523,6 +1523,7 @@ import Mathlib.CategoryTheory.ChosenFiniteProducts import Mathlib.CategoryTheory.ChosenFiniteProducts.Cat import Mathlib.CategoryTheory.ChosenFiniteProducts.FunctorCategory import Mathlib.CategoryTheory.Closed.Cartesian +import Mathlib.CategoryTheory.Closed.Enrichment import Mathlib.CategoryTheory.Closed.Functor import Mathlib.CategoryTheory.Closed.FunctorCategory import Mathlib.CategoryTheory.Closed.FunctorToTypes diff --git a/Mathlib/CategoryTheory/Closed/Enrichment.lean b/Mathlib/CategoryTheory/Closed/Enrichment.lean new file mode 100644 index 0000000000000..1d263f92528ed --- /dev/null +++ b/Mathlib/CategoryTheory/Closed/Enrichment.lean @@ -0,0 +1,41 @@ +/- +Copyright (c) 2024 Daniel Carranza. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Daniel Carranza +-/ +import Mathlib.CategoryTheory.Enriched.Basic +import Mathlib.CategoryTheory.Closed.Monoidal + +/-! +# A closed monoidal category is enriched in itself + +From the data of a closed monoidal category `C`, we define a `C`-category structure for `C`. +where the hom-object is given by the internal hom (coming from the closed structure). + +We use `scoped instance` to avoid potential issues where `C` may also have +a `C`-category structure coming from another source (e.g. the type of simplicial sets +`SSet.{v}` has an instance of `EnrichedCategory SSet.{v}` as a category of simplicial objects; +see `AlgebraicTopology/SimplicialCategory/SimplicialObject`). + +All structure field values are defined in `Closed/Monoidal`. + +-/ + +universe u v + +namespace CategoryTheory + +namespace MonoidalClosed + +variable {C : Type u} [Category.{v} C] [MonoidalCategory C] [MonoidalClosed C] + +/-- For `C` closed monoidal, build an instance of `C` as a `C`-category -/ +scoped instance : EnrichedCategory C C where + Hom x := (ihom x).obj + id _ := id _ + comp _ _ _ := comp _ _ _ + assoc _ _ _ _ := assoc _ _ _ _ + +end MonoidalClosed + +end CategoryTheory diff --git a/Mathlib/CategoryTheory/Closed/Monoidal.lean b/Mathlib/CategoryTheory/Closed/Monoidal.lean index 43af4d461c0de..fd270e7d51120 100644 --- a/Mathlib/CategoryTheory/Closed/Monoidal.lean +++ b/Mathlib/CategoryTheory/Closed/Monoidal.lean @@ -304,6 +304,89 @@ theorem ofEquiv_uncurry_def {X Y Z : C} : end OfEquiv +-- A closed monoidal category C is always enriched over itself. +-- This section contains the necessary definitions and equalities to endow C with +-- the structure of a C-category, while the instance itself is defined in `Closed/Enrichment`. +-- In particular, we only assume the necessary instances of `Closed x`, rather than assuming +-- C comes with an instance of `MonoidalClosed` +section Enriched + +/-- The C-identity morphism + `𝟙_ C ⟶ hom(x, x)` +used to equip `C` with the structure of a `C`-category -/ +def id (x : C) [Closed x] : 𝟙_ C ⟶ (ihom x).obj x := curry (ρ_ x).hom + +/-- The *uncurried* composition morphism + `x ⊗ (hom(x, y) ⊗ hom(y, z)) ⟶ (x ⊗ hom(x, y)) ⊗ hom(y, z) ⟶ y ⊗ hom(y, z) ⟶ z`. +The `C`-composition morphism will be defined as the adjoint transpose of this map. -/ +def compTranspose (x y z : C) [Closed x] [Closed y] : x ⊗ (ihom x).obj y ⊗ (ihom y).obj z ⟶ z := + (α_ x ((ihom x).obj y) ((ihom y).obj z)).inv ≫ + (ihom.ev x).app y ▷ ((ihom y).obj z) ≫ (ihom.ev y).app z + +/-- The `C`-composition morphism + `hom(x, y) ⊗ hom(y, z) ⟶ hom(x, z)` +used to equip `C` with the structure of a `C`-category -/ +def comp (x y z : C) [Closed x] [Closed y] : (ihom x).obj y ⊗ (ihom y).obj z ⟶ (ihom x).obj z := + curry (compTranspose x y z) + +/-- Unfold the definition of `id`. +This exists to streamline the proofs of `MonoidalClosed.id_comp` and `MonoidalClosed.comp_id` -/ +lemma id_eq (x : C) [Closed x] : id x = curry (ρ_ x).hom := rfl + +/-- Unfold the definition of `compTranspose`. +This exists to streamline the proof of `MonoidalClosed.assoc` -/ +lemma compTranspose_eq (x y z : C) [Closed x] [Closed y] : + compTranspose x y z = (α_ _ _ _).inv ≫ (ihom.ev x).app y ▷ _ ≫ (ihom.ev y).app z := + rfl + +/-- Unfold the definition of `comp`. +This exists to streamline the proof of `MonoidalClosed.assoc` -/ +lemma comp_eq (x y z : C) [Closed x] [Closed y] : comp x y z = curry (compTranspose x y z) := rfl + +/-! +The proofs of associativity and unitality use the following outline: + 1. Take adjoint transpose on each side of the equality (uncurry_injective) + 2. Do whatever rewrites/simps are necessary to apply uncurry_curry + 3. Conclude with simp +-/ + +/-- Left unitality of the enriched structure -/ +@[reassoc (attr := simp)] +lemma id_comp (x y : C) [Closed x] : + (λ_ ((ihom x).obj y)).inv ≫ id x ▷ _ ≫ comp x x y = 𝟙 _:= by + apply uncurry_injective + rw [uncurry_natural_left, uncurry_natural_left, comp_eq, uncurry_curry, id_eq, compTranspose_eq, + associator_inv_naturality_middle_assoc, ← comp_whiskerRight_assoc, ← uncurry_eq, + uncurry_curry, triangle_assoc_comp_right_assoc, whiskerLeft_inv_hom_assoc, + uncurry_id_eq_ev _ _] + +/-- Right unitality of the enriched structure -/ +@[reassoc (attr := simp)] +lemma comp_id (x y : C) [Closed x] [Closed y] : + (ρ_ ((ihom x).obj y)).inv ≫ _ ◁ id y ≫ comp x y y = 𝟙 _ := by + apply uncurry_injective + rw [uncurry_natural_left, uncurry_natural_left, comp_eq, uncurry_curry, compTranspose_eq, + associator_inv_naturality_right_assoc, ← rightUnitor_tensor_inv_assoc, + whisker_exchange_assoc, ← rightUnitor_inv_naturality_assoc, ← uncurry_id_eq_ev y y] + simp only [Functor.id_obj] + rw [← uncurry_natural_left] + simp [id_eq, uncurry_id_eq_ev] + +/-- Associativity of the enriched structure -/ +@[reassoc] +lemma assoc (w x y z : C) [Closed w] [Closed x] [Closed y] : + (α_ _ _ _).inv ≫ comp w x y ▷ _ ≫ comp w y z = _ ◁ comp x y z ≫ comp w x z := by + apply uncurry_injective + simp only [uncurry_natural_left, comp_eq] + rw [uncurry_curry, uncurry_curry]; simp only [compTranspose_eq, Category.assoc] + rw [associator_inv_naturality_middle_assoc, ← comp_whiskerRight_assoc]; dsimp + rw [← uncurry_eq, uncurry_curry, associator_inv_naturality_right_assoc, whisker_exchange_assoc, + ← uncurry_eq, uncurry_curry] + simp only [comp_whiskerRight, tensorLeft_obj, Category.assoc, pentagon_inv_assoc, + whiskerRight_tensor, Iso.hom_inv_id_assoc] + +end Enriched + end MonoidalClosed attribute [nolint simpNF] CategoryTheory.MonoidalClosed.homEquiv_apply_eq CategoryTheory.MonoidalClosed.homEquiv_symm_apply_eq From 9dda7f5db42f6c0e303ca618e1b21995ac435f2f Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Thu, 24 Oct 2024 22:28:20 +0000 Subject: [PATCH 397/505] chore(create-adaptation-pr.sh): more escaping in strings (#18108) --- scripts/create-adaptation-pr.sh | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/scripts/create-adaptation-pr.sh b/scripts/create-adaptation-pr.sh index e383694363dac..8af305216c0e8 100755 --- a/scripts/create-adaptation-pr.sh +++ b/scripts/create-adaptation-pr.sh @@ -204,7 +204,7 @@ if git diff --name-only bump/$BUMPVERSION bump/nightly-$NIGHTLYDATE | grep -q .; echo "### [auto/user] post a link to the PR on Zulip" zulip_title="#$pr_number adaptations for nightly-$NIGHTLYDATE" - zulip_body="> $pr_title #$pr_number\n\nPlease review this PR. At the end of the month this diff will land in \`master\`." + zulip_body="> $pr_title #$pr_number\\\n\\\nPlease review this PR. At the end of the month this diff will land in \\\`master\\\`." echo "Post the link to the PR in a new thread on the #nightly-testing channel on Zulip" echo "Here is a suggested message:" From 62764f8ffe78da82e4e48925dd51aa017acb62a4 Mon Sep 17 00:00:00 2001 From: grunweg Date: Thu, 24 Oct 2024 22:41:18 +0000 Subject: [PATCH 398/505] feat: extend the linter against broad imports (#14898) Also lint on importing `Tactic.{Have,Replace}` or any file in the `Deprecated` directory. Both of these are not supposed to be imported; this linter enforces it. --- Mathlib/Tactic/Linter/Header.lean | 18 ++++++++++++++++-- test/Header.lean | 8 ++++++++ 2 files changed, 24 insertions(+), 2 deletions(-) diff --git a/Mathlib/Tactic/Linter/Header.lean b/Mathlib/Tactic/Linter/Header.lean index bcc44a5ae56ad..4089b703bfc7e 100644 --- a/Mathlib/Tactic/Linter/Header.lean +++ b/Mathlib/Tactic/Linter/Header.lean @@ -207,18 +207,32 @@ def copyrightHeaderChecks (copyright : String) : Array (Syntax × String) := Id. /-- Check the `Syntax` `imports` for broad imports: either `Mathlib.Tactic` or any import starting with `Lake`. -/ -def broadImportsCheck (imports : Array Syntax) : Array (Syntax × String) := Id.run do +def broadImportsCheck (imports : Array Syntax) (mainModule : Name) : Array (Syntax × String) := Id.run do let mut output := #[] for i in imports do match i.getId with | `Mathlib.Tactic => output := output.push (i, s!"Files in mathlib cannot import the whole tactic folder.") + | `Mathlib.Tactic.Replace => + if mainModule != `Mathlib.Tactic then output := output.push (i, + s!"Mathlib.Tactic.Replace defines a deprecated form of the 'replace' tactic; \ + please do not use it in mathlib.") + | `Mathlib.Tactic.Have => + if ![`Mathlib.Tactic, `Mathlib.Tactic.Replace].contains mainModule then + output := output.push (i, + s!"Mathlib.Tactic.Have defines a deprecated form of the 'have' tactic; \ + please do not use it in mathlib.") | modName => if modName.getRoot == `Lake then output := output.push (i, s!"In the past, importing 'Lake' in mathlib has led to dramatic slow-downs of the linter \ (see e.g. mathlib4#13779). Please consider carefully if this import is useful and \ make sure to benchmark it. If this is fine, feel free to allow this linter.") + else if (`Mathlib.Deprecated).isPrefixOf modName && + !(`Mathlib.Deprecated).isPrefixOf mainModule then + -- We do not complain about files in the `Deprecated` directory importing one another. + output := output.push (i, s!"Files in the `Deprecated` directory are not supposed to be imported.") + return output /-- @@ -277,7 +291,7 @@ def headerLinter : Linter where run := withSetOptionIn fun stx ↦ do parseUpToHere (stx.getTailPos?.getD default) "\nsection") let importIds := getImportIds upToStx -- Report on broad imports. - for (imp, msg) in broadImportsCheck importIds do + for (imp, msg) in broadImportsCheck importIds mainModule do Linter.logLint linter.style.header imp msg let afterImports := firstNonImport? upToStx if afterImports.isNone then return diff --git a/test/Header.lean b/test/Header.lean index f1b9352eb9404..5644cae930fee 100644 --- a/test/Header.lean +++ b/test/Header.lean @@ -7,6 +7,8 @@ Authors: Damiano Testa import Mathlib.Tactic.Linter.Header import Lake import /- -/ Mathlib.Tactic -- the `TextBased` linter does not flag this `broadImport` +import Mathlib.Tactic.Have +import Mathlib.Deprecated.Subfield /-- warning: In the past, importing 'Lake' in mathlib has led to dramatic slow-downs of the linter (see e.g. mathlib4#13779). Please consider carefully if this import is useful and make sure to benchmark it. If this is fine, feel free to allow this linter. @@ -15,6 +17,12 @@ note: this linter can be disabled with `set_option linter.style.header false` warning: Files in mathlib cannot import the whole tactic folder. note: this linter can be disabled with `set_option linter.style.header false` --- +warning: Mathlib.Tactic.Have defines a deprecated form of the 'have' tactic; please do not use it in mathlib. +note: this linter can be disabled with `set_option linter.style.header false` +--- +warning: Files in the `Deprecated` directory are not supposed to be imported. +note: this linter can be disabled with `set_option linter.style.header false` +--- warning: The module doc-string for a file should be the first command after the imports. Please, add a module doc-string before `/-!# Tests for the `docModule` linter -/ From a1578ecc8f1ed2f7bef68f3ff7c321bfc0c49820 Mon Sep 17 00:00:00 2001 From: Matthew Robert Ballard Date: Fri, 25 Oct 2024 00:50:18 +0000 Subject: [PATCH 399/505] chore: use `Batteries` test driver directly in the lake file (#15897) With leanprover/lean4#4261, we can specify a test driver from a dependency. We do this with `Batteries` test driver. Previously, the test driver from `Batteries` was called via `scripts/test.lean` which is now removed. --- lakefile.lean | 14 ++------------ scripts/test.lean | 21 --------------------- 2 files changed, 2 insertions(+), 33 deletions(-) delete mode 100644 scripts/test.lean diff --git a/lakefile.lean b/lakefile.lean index 697f41bab4e5b..25b17ebe85a33 100644 --- a/lakefile.lean +++ b/lakefile.lean @@ -52,6 +52,8 @@ package mathlib where leanOptions := mathlibLeanOptions -- Mathlib also enforces these linter options, which are not active by default. moreServerOptions := mathlibOnlyLinters + -- Use Batteries' test driver for `lake test` + testDriver := "batteries/test" -- These are additional settings which do not affect the lake hash, -- so they can be enabled in CI and disabled locally or vice versa. -- Warning: Do not put any options here that actually change the olean files, @@ -132,18 +134,6 @@ lean_exe pole where -- Executables which import `Lake` must set `-lLake`. weakLinkArgs := #["-lLake"] -/-- -`lake exe test` is a thin wrapper around `lake exe batteries/test`, until -https://github.com/leanprover/lean4/issues/4121 is resolved. - -You can also use it as e.g. `lake exe test conv eval_elab` to only run the named tests. --/ -@[test_driver] -lean_exe test where - -- We could add the above `leanOptions` and `moreServerOptions`: currently, these do not take - -- effect as `test` is a `lean_exe`. With a `lean_lib`, it would work... - srcDir := "scripts" - /-! ## Other configuration -/ diff --git a/scripts/test.lean b/scripts/test.lean deleted file mode 100644 index f02a621637377..0000000000000 --- a/scripts/test.lean +++ /dev/null @@ -1,21 +0,0 @@ -/- -Copyright (c) 2024 Kim Morrison. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Kim Morrison --/ -open IO.Process -open System - -/-- -Run tests, via the Batteries test runner. - -When https://github.com/leanprover/lean4/issues/4121 -"allow using an upstream executable as a lake `@[test_runner]`" -is resolved, this file can be replaced with a line in `lakefile.lean`. --/ -def main (args : List String) : IO Unit := do - -- ProofWidgets and Batteries may not have been completely built by `lake build` yet, - -- but they are needed by some tests. - _ ← (← spawn { cmd := "lake", args := #["build", "ProofWidgets", "Batteries"] }).wait - let exitcode ← (← spawn { cmd := "lake", args := #["exe", "batteries/test"] ++ args }).wait - exit exitcode.toUInt8 From 87ee1f25bf4ee48b5290edf466511e6e76d7286a Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Fri, 25 Oct 2024 01:40:44 +0000 Subject: [PATCH 400/505] chore: cleanup of `mergeSort'` (#16902) This PR removes the old `mergeSort'` in Mathlib (which was *not* a stable sort algorithm), replacing all uses with the stable `mergeSort` from Lean, mostly preserving the API. --- Mathlib/Data/List/Sort.lean | 239 +++++++------------------------- Mathlib/Data/Multiset/Sort.lean | 30 ++-- Mathlib/Logic/Equiv/List.lean | 4 +- Mathlib/Order/Defs.lean | 32 +++++ 4 files changed, 102 insertions(+), 203 deletions(-) diff --git a/Mathlib/Data/List/Sort.lean b/Mathlib/Data/List/Sort.lean index 513e81df6f0a3..6733b9df5e8a3 100644 --- a/Mathlib/Data/List/Sort.lean +++ b/Mathlib/Data/List/Sort.lean @@ -3,39 +3,20 @@ Copyright (c) 2016 Jeremy Avigad. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad -/ +import Batteries.Data.List.Pairwise import Mathlib.Data.List.OfFn import Mathlib.Data.List.Nodup import Mathlib.Order.Fin.Basic -import Batteries.Data.List.Perm /-! # Sorting algorithms on lists In this file we define `List.Sorted r l` to be an alias for `List.Pairwise r l`. This alias is preferred in the case that `r` is a `<` or `≤`-like relation. -Then we define two sorting algorithms: -`List.insertionSort` and `List.mergeSort'`, and prove their correctness. +Then we define the sorting algorithm +`List.insertionSort` and prove its correctness. -/ -#adaptation_note -/-- -`List.mergeSort` has now been implemented in Lean4. -It improves on the one here by being a "stable" sort -(in the sense that a sorted sublist of the original list remains a sublist of the result), -and is also marginally faster. - -However we haven't yet replaced `List.mergeSort'` here. -The obstacle is that `mergeSort'` is written using `r : α → α → Prop` with `[DecidableRel r]`, -while `mergeSort` uses `r : α → α → Bool`. This is hardly insurmountable, -but it's a bit of work that hasn't been done yet. - -`List.mergeSort'` is only used in Mathlib to sort multisets for printing, so this is not critical. - -A pull request cleaning up here, and ideally deprecating or deleting `List.mergeSort'`, -would be welcome. --/ - - open List.Perm universe u v @@ -165,6 +146,14 @@ theorem Sorted.rel_of_mem_take_of_mem_drop {l : List α} (h : List.Sorted r l) { rw [length_take] at hix exact h.rel_get_of_lt (Nat.lt_add_right _ (Nat.lt_min.mp hix).left) +/-- +If a list is sorted with respect to a decidable relation, +then it is sorted with respect to the corresponding Bool-valued relation. +-/ +theorem Sorted.decide [DecidableRel r] (l : List α) (h : Sorted r l) : + Sorted (fun a b => decide (r a b) = true) l := by + refine h.imp fun {a b} h => by simpa using h + end Sorted section Monotone @@ -220,6 +209,11 @@ def insertionSort : List α → List α | [] => [] | b :: l => orderedInsert r b (insertionSort l) +-- A quick check that insertionSort is stable: +example : + insertionSort (fun m n => m / 10 ≤ n / 10) [5, 27, 221, 95, 17, 43, 7, 2, 98, 567, 23, 12] = + [5, 7, 2, 17, 12, 27, 23, 43, 95, 98, 221, 567] := rfl + @[simp] theorem orderedInsert_nil (a : α) : [].orderedInsert r a = [a] := rfl @@ -492,192 +486,61 @@ end Correctness end InsertionSort -/-! ### Merge sort -/ +/-! ### Merge sort +We provide some wrapper functions around the theorems for `mergeSort` provided in Lean, +which rather than using explicit hypotheses for transitivity and totality, +use Mathlib order typeclasses instead. +-/ -section MergeSort - --- TODO(Jeremy): observation: if instead we write (a :: (split l).1, b :: (split l).2), the --- equation compiler can't prove the third equation -/-- Split `l` into two lists of approximately equal length. - - split [1, 2, 3, 4, 5] = ([1, 3, 5], [2, 4]) -/ -@[simp] -def split : List α → List α × List α - | [] => ([], []) - | a :: l => - let (l₁, l₂) := split l - (a :: l₂, l₁) - -theorem split_cons_of_eq (a : α) {l l₁ l₂ : List α} (h : split l = (l₁, l₂)) : - split (a :: l) = (a :: l₂, l₁) := by rw [split, h] +unseal merge mergeSort in +example : + mergeSort [5, 27, 221, 95, 17, 43, 7, 2, 98, 567, 23, 12] (fun m n => m / 10 ≤ n / 10) = + [5, 7, 2, 17, 12, 27, 23, 43, 95, 98, 221, 567] := rfl -@[simp] -theorem map_split (f : α → β) : - ∀ l : List α, (map f l).split = (l.split.1.map f, l.split.2.map f) - | [] => rfl - | a :: l => by simp [map_split] - -@[simp] -theorem mem_split_iff {x : α} : ∀ {l : List α}, x ∈ l.split.1 ∨ x ∈ l.split.2 ↔ x ∈ l - | [] => by simp - | a :: l => by simp_rw [split, mem_cons, or_assoc, or_comm, mem_split_iff] - -theorem length_split_le : - ∀ {l l₁ l₂ : List α}, split l = (l₁, l₂) → length l₁ ≤ length l ∧ length l₂ ≤ length l - | [], _, _, rfl => ⟨Nat.le_refl 0, Nat.le_refl 0⟩ - | a :: l, l₁', l₂', h => by - cases' e : split l with l₁ l₂ - injection (split_cons_of_eq _ e).symm.trans h; substs l₁' l₂' - cases' length_split_le e with h₁ h₂ - exact ⟨Nat.succ_le_succ h₂, Nat.le_succ_of_le h₁⟩ - -theorem length_split_fst_le (l : List α) : length (split l).1 ≤ length l := - (length_split_le rfl).1 - -theorem length_split_snd_le (l : List α) : length (split l).2 ≤ length l := - (length_split_le rfl).2 - -theorem length_split_lt {a b} {l l₁ l₂ : List α} (h : split (a :: b :: l) = (l₁, l₂)) : - length l₁ < length (a :: b :: l) ∧ length l₂ < length (a :: b :: l) := by - cases' e : split l with l₁' l₂' - injection (split_cons_of_eq _ (split_cons_of_eq _ e)).symm.trans h; substs l₁ l₂ - cases' length_split_le e with h₁ h₂ - exact ⟨Nat.succ_le_succ (Nat.succ_le_succ h₁), Nat.succ_le_succ (Nat.succ_le_succ h₂)⟩ - -theorem perm_split : ∀ {l l₁ l₂ : List α}, split l = (l₁, l₂) → l ~ l₁ ++ l₂ - | [], _, _, rfl => Perm.refl _ - | a :: l, l₁', l₂', h => by - cases' e : split l with l₁ l₂ - injection (split_cons_of_eq _ e).symm.trans h; substs l₁' l₂' - exact ((perm_split e).trans perm_append_comm).cons a - -/-- Implementation of a merge sort algorithm to sort a list. -/ -def mergeSort' : List α → List α - | [] => [] - | [a] => [a] - | a :: b :: l => by - -- Porting note: rewrote to make `mergeSort_cons_cons` proof easier - let ls := (split (a :: b :: l)) - have := length_split_fst_le l - have := length_split_snd_le l - exact merge (mergeSort' ls.1) (mergeSort' ls.2) (r · ·) - termination_by l => length l - -@[nolint unusedHavesSuffices] -- Porting note: false positive -theorem mergeSort'_cons_cons {a b} {l l₁ l₂ : List α} (h : split (a :: b :: l) = (l₁, l₂)) : - mergeSort' r (a :: b :: l) = merge (mergeSort' r l₁) (mergeSort' r l₂) (r · ·) := by - simp only [mergeSort', h] +section MergeSort section Correctness -theorem perm_mergeSort' : ∀ l : List α, mergeSort' r l ~ l - | [] => by simp [mergeSort'] - | [a] => by simp [mergeSort'] - | a :: b :: l => by - cases' e : split (a :: b :: l) with l₁ l₂ - cases' length_split_lt e with h₁ h₂ - rw [mergeSort'_cons_cons r e] - apply (perm_merge (r · ·) _ _).trans - exact - ((perm_mergeSort' l₁).append (perm_mergeSort' l₂)).trans (perm_split e).symm - termination_by l => length l - -@[simp] -theorem mem_mergeSort' {l : List α} {x : α} : x ∈ l.mergeSort' r ↔ x ∈ l := - (perm_mergeSort' r l).mem_iff - -@[simp] -theorem length_mergeSort' (l : List α) : (mergeSort' r l).length = l.length := - (perm_mergeSort' r _).length_eq - section TotalAndTransitive variable {r} [IsTotal α r] [IsTrans α r] -theorem Sorted.merge : ∀ {l l' : List α}, Sorted r l → Sorted r l' → Sorted r (merge l l' (r · ·) ) - | [], [], _, _ => by simp - | [], b :: l', _, h₂ => by simpa using h₂ - | a :: l, [], h₁, _ => by simpa using h₁ - | a :: l, b :: l', h₁, h₂ => by - by_cases h : a ≼ b - · suffices ∀ b' ∈ List.merge l (b :: l') (r · ·) , r a b' by - simpa [h, h₁.of_cons.merge h₂] - intro b' bm - rcases show b' = b ∨ b' ∈ l ∨ b' ∈ l' by - simpa [or_left_comm] using (perm_merge _ _ _).subset bm with - (be | bl | bl') - · subst b' - assumption - · exact rel_of_sorted_cons h₁ _ bl - · exact _root_.trans h (rel_of_sorted_cons h₂ _ bl') - · suffices ∀ b' ∈ List.merge (a :: l) l' (r · ·) , r b b' by - simpa [h, h₁.merge h₂.of_cons] - intro b' bm - have ba : b ≼ a := (total_of r _ _).resolve_left h - have : b' = a ∨ b' ∈ l ∨ b' ∈ l' := by simpa using (perm_merge _ _ _).subset bm - rcases this with (be | bl | bl') - · subst b' - assumption - · exact _root_.trans ba (rel_of_sorted_cons h₁ _ bl) - · exact rel_of_sorted_cons h₂ _ bl' +theorem Sorted.merge {l l' : List α} (h : Sorted r l) (h' : Sorted r l') : + Sorted r (merge l l' (r · ·)) := by + simpa using sorted_merge (le := (r · ·)) + (fun a b c h₁ h₂ => by simpa using _root_.trans (by simpa using h₁) (by simpa using h₂)) + (fun a b => by simpa using IsTotal.total a b) + l l' (by simpa using h) (by simpa using h') variable (r) -theorem sorted_mergeSort' : ∀ l : List α, Sorted r (mergeSort' r l) - | [] => by simp [mergeSort'] - | [a] => by simp [mergeSort'] - | a :: b :: l => by - cases' e : split (a :: b :: l) with l₁ l₂ - cases' length_split_lt e with h₁ h₂ - rw [mergeSort'_cons_cons r e] - exact (sorted_mergeSort' l₁).merge (sorted_mergeSort' l₂) - termination_by l => length l - -theorem mergeSort'_eq_self [IsAntisymm α r] {l : List α} : Sorted r l → mergeSort' r l = l := - eq_of_perm_of_sorted (perm_mergeSort' _ _) (sorted_mergeSort' _ _) - -theorem mergeSort'_eq_insertionSort [IsAntisymm α r] (l : List α) : - mergeSort' r l = insertionSort r l := - eq_of_perm_of_sorted ((perm_mergeSort' r l).trans (perm_insertionSort r l).symm) - (sorted_mergeSort' r l) (sorted_insertionSort r l) +/-- Variant of `sorted_mergeSort` using order typeclasses. -/ +theorem sorted_mergeSort' [Preorder α] [DecidableRel ((· : α) ≤ ·)] [IsTotal α (· ≤ ·)] + (l : List α) : Sorted (· ≤ ·) (mergeSort l) := by + simpa using sorted_mergeSort (le := fun a b => a ≤ b) + (fun a b c h₁ h₂ => by simpa using le_trans (by simpa using h₁) (by simpa using h₂)) + (fun a b => by simpa using IsTotal.total a b) + l + +theorem mergeSort_eq_self [LinearOrder α] {l : List α} : Sorted (· ≤ ·) l → mergeSort l = l := + eq_of_perm_of_sorted (mergeSort_perm _ _) (sorted_mergeSort' l) + +theorem mergeSort_eq_insertionSort [IsAntisymm α r] (l : List α) : + mergeSort l (r · ·) = insertionSort r l := + eq_of_perm_of_sorted ((mergeSort_perm l _).trans (perm_insertionSort r l).symm) + (sorted_mergeSort (le := (r · ·)) + (fun a b c h₁ h₂ => by simpa using _root_.trans (by simpa using h₁) (by simpa using h₂)) + (fun a b => by simpa using IsTotal.total a b) + l) + (sorted_insertionSort r l).decide end TotalAndTransitive end Correctness -@[simp] -theorem mergeSort'_nil : [].mergeSort' r = [] := by rw [List.mergeSort'] - -@[simp] -theorem mergeSort'_singleton (a : α) : [a].mergeSort' r = [a] := by rw [List.mergeSort'] - -theorem map_mergeSort' (f : α → β) (l : List α) (hl : ∀ a ∈ l, ∀ b ∈ l, a ≼ b ↔ f a ≼ f b) : - (l.mergeSort' r).map f = (l.map f).mergeSort' s := - match l with - | [] => by simp - | [x] => by simp - | a :: b :: l => by - simp_rw [← mem_split_iff (l := a :: b :: l), or_imp, forall_and] at hl - set l₁ := (split (a :: b :: l)).1 - set l₂ := (split (a :: b :: l)).2 - have e : split (a :: b :: l) = (l₁, l₂) := rfl - have fe : split (f a :: f b :: l.map f) = (l₁.map f, l₂.map f) := by - rw [← map, ← map, map_split, e] - have := length_split_fst_le l - have := length_split_snd_le l - simp_rw [List.map] - rw [List.mergeSort'_cons_cons _ e, List.mergeSort'_cons_cons _ fe, - map_merge, map_mergeSort' _ l₁ hl.1.1, map_mergeSort' _ l₂ hl.2.2] - simp_rw [mem_mergeSort', decide_eq_decide] - exact hl.1.2 - termination_by length l - end MergeSort end sort --- try them out! ---#eval insertionSort (fun m n : ℕ => m ≤ n) [5, 27, 221, 95, 17, 43, 7, 2, 98, 567, 23, 12] ---#eval mergeSort' (fun m n : ℕ => m ≤ n) [5, 27, 221, 95, 17, 43, 7, 2, 98, 567, 23, 12] end List diff --git a/Mathlib/Data/Multiset/Sort.lean b/Mathlib/Data/Multiset/Sort.lean index 997af93b548c0..7c329e6f85d53 100644 --- a/Mathlib/Data/Multiset/Sort.lean +++ b/Mathlib/Data/Multiset/Sort.lean @@ -10,13 +10,12 @@ import Mathlib.Data.Multiset.Basic # Construct a sorted list from a multiset. -/ +variable {α β : Type*} namespace Multiset open List -variable {α β : Type*} - section sort variable (r : α → α → Prop) [DecidableRel r] [IsTrans α r] [IsAntisymm α r] [IsTotal α r] @@ -25,47 +24,52 @@ variable (r' : β → β → Prop) [DecidableRel r'] [IsTrans β r'] [IsAntisymm /-- `sort s` constructs a sorted list from the multiset `s`. (Uses merge sort algorithm.) -/ def sort (s : Multiset α) : List α := - Quot.liftOn s (mergeSort' r) fun _ _ h => - eq_of_perm_of_sorted ((perm_mergeSort' _ _).trans <| h.trans (perm_mergeSort' _ _).symm) - (sorted_mergeSort' r _) (sorted_mergeSort' r _) + Quot.liftOn s (mergeSort · (r · ·)) fun _ _ h => + eq_of_perm_of_sorted ((mergeSort_perm _ _).trans <| h.trans (mergeSort_perm _ _).symm) + (sorted_mergeSort IsTrans.trans + (fun a b => by simpa using IsTotal.total a b) _) + (sorted_mergeSort IsTrans.trans + (fun a b => by simpa using IsTotal.total a b) _) @[simp] -theorem coe_sort (l : List α) : sort r l = mergeSort' r l := +theorem coe_sort (l : List α) : sort r l = mergeSort l (r · ·) := rfl @[simp] theorem sort_sorted (s : Multiset α) : Sorted r (sort r s) := - Quot.inductionOn s fun _l => sorted_mergeSort' r _ + Quot.inductionOn s fun l => by + simpa using sorted_mergeSort (le := (r · ·)) IsTrans.trans + (fun a b => by simpa using IsTotal.total a b) l @[simp] theorem sort_eq (s : Multiset α) : ↑(sort r s) = s := - Quot.inductionOn s fun _ => Quot.sound <| perm_mergeSort' _ _ + Quot.inductionOn s fun _ => Quot.sound <| mergeSort_perm _ _ @[simp] theorem mem_sort {s : Multiset α} {a : α} : a ∈ sort r s ↔ a ∈ s := by rw [← mem_coe, sort_eq] @[simp] theorem length_sort {s : Multiset α} : (sort r s).length = card s := - Quot.inductionOn s <| length_mergeSort' _ + Quot.inductionOn s <| length_mergeSort @[simp] theorem sort_zero : sort r 0 = [] := - List.mergeSort'_nil r + List.mergeSort_nil @[simp] theorem sort_singleton (a : α) : sort r {a} = [a] := - List.mergeSort'_singleton r a + List.mergeSort_singleton a theorem map_sort (f : α → β) (s : Multiset α) (hs : ∀ a ∈ s, ∀ b ∈ s, r a b ↔ r' (f a) (f b)) : (s.sort r).map f = (s.map f).sort r' := by revert s - exact Quot.ind fun _ => List.map_mergeSort' _ _ _ _ + exact Quot.ind fun l h => map_mergeSort (l := l) (by simpa using h) theorem sort_cons (a : α) (s : Multiset α) : (∀ b ∈ s, r a b) → sort r (a ::ₘ s) = a :: sort r s := by refine Quot.inductionOn s fun l => ?_ - simpa [mergeSort'_eq_insertionSort] using insertionSort_cons r + simpa [mergeSort_eq_insertionSort] using insertionSort_cons r (a := a) (l := l) end sort diff --git a/Mathlib/Logic/Equiv/List.lean b/Mathlib/Logic/Equiv/List.lean index bf14924bbdaac..c0753898bf0dc 100644 --- a/Mathlib/Logic/Equiv/List.lean +++ b/Mathlib/Logic/Equiv/List.lean @@ -289,7 +289,7 @@ instance multiset : Denumerable (Multiset α) := raise_lower (List.sorted_cons.2 ⟨fun n _ => Nat.zero_le n, (s.map encode).sort_sorted _⟩) simp [-Multiset.map_coe, this], fun n => by - simp [-Multiset.map_coe, List.mergeSort'_eq_self _ (raise_sorted _ _), lower_raise]⟩ + simp [-Multiset.map_coe, List.mergeSort_eq_self (raise_sorted _ _), lower_raise]⟩ end Multiset @@ -344,7 +344,7 @@ instance finset : Denumerable (Finset α) := raise_lower' (fun n _ => Nat.zero_le n) (Finset.sort_sorted_lt _)], fun n => by simp [-Multiset.map_coe, Finset.map, raise'Finset, Finset.sort, - List.mergeSort'_eq_self (· ≤ ·) ((raise'_sorted _ _).imp (@le_of_lt _ _)), lower_raise']⟩ + List.mergeSort_eq_self ((raise'_sorted _ _).imp (@le_of_lt _ _)), lower_raise']⟩ end Finset diff --git a/Mathlib/Order/Defs.lean b/Mathlib/Order/Defs.lean index 75d603ee2c080..339a1a55bc927 100644 --- a/Mathlib/Order/Defs.lean +++ b/Mathlib/Order/Defs.lean @@ -125,6 +125,38 @@ instance (priority := 90) isAsymm_of_isTrans_of_isIrrefl [IsTrans α r] [IsIrref IsAsymm α r := ⟨fun a _b h₁ h₂ => absurd (_root_.trans h₁ h₂) (irrefl a)⟩ +instance IsIrrefl.decide [DecidableRel r] [IsIrrefl α r] : + IsIrrefl α (fun a b => decide (r a b) = true) where + irrefl := fun a => by simpa using irrefl a + +instance IsRefl.decide [DecidableRel r] [IsRefl α r] : + IsRefl α (fun a b => decide (r a b) = true) where + refl := fun a => by simpa using refl a + +instance IsTrans.decide [DecidableRel r] [IsTrans α r] : + IsTrans α (fun a b => decide (r a b) = true) where + trans := fun a b c => by simpa using trans a b c + +instance IsSymm.decide [DecidableRel r] [IsSymm α r] : + IsSymm α (fun a b => decide (r a b) = true) where + symm := fun a b => by simpa using symm a b + +instance IsAntisymm.decide [DecidableRel r] [IsAntisymm α r] : + IsAntisymm α (fun a b => decide (r a b) = true) where + antisymm := fun a b h₁ h₂ => antisymm _ _ (by simpa using h₁) (by simpa using h₂) + +instance IsAsymm.decide [DecidableRel r] [IsAsymm α r] : + IsAsymm α (fun a b => decide (r a b) = true) where + asymm := fun a b => by simpa using asymm a b + +instance IsTotal.decide [DecidableRel r] [IsTotal α r] : + IsTotal α (fun a b => decide (r a b) = true) where + total := fun a b => by simpa using total a b + +instance IsTrichotomous.decide [DecidableRel r] [IsTrichotomous α r] : + IsTrichotomous α (fun a b => decide (r a b) = true) where + trichotomous := fun a b => by simpa using trichotomous a b + variable (r) @[elab_without_expected_type] lemma irrefl_of [IsIrrefl α r] (a : α) : ¬a ≺ a := irrefl a From fc157081b3adaa6877bf68c0415d54c0126cf47f Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Fri, 25 Oct 2024 05:50:35 +0000 Subject: [PATCH 401/505] chore(Algebra/GroupWithZero/Action): reduce theory included with `Group.Action.Defs` (#18190) The `GroupWithZero/Action/Defs.lean` file depends on a bit of theory and in turn defines some results not needed for e.g. modules. This PR tries to ensure that the focus is on the definitions of `SMulZeroClass` and `DistribMulAction`. In the process, I made the following moves: * `Action/End.lean` now contains results involving `Function.End` and some results relating homomorphisms with multiplication * `Action/Faithful.lean` now contains results involving `FaithfulAction` --- Mathlib.lean | 2 + .../Group/Submonoid/DistribMulAction.lean | 2 +- .../Algebra/GroupWithZero/Action/Defs.lean | 116 +----------------- Mathlib/Algebra/GroupWithZero/Action/End.lean | 93 ++++++++++++++ .../GroupWithZero/Action/Faithful.lean | 23 ++++ .../Algebra/GroupWithZero/Action/Prod.lean | 2 +- .../Algebra/GroupWithZero/Action/Units.lean | 35 +++++- Mathlib/Algebra/Module/Defs.lean | 2 + Mathlib/Algebra/Ring/Action/Basic.lean | 2 +- Mathlib/Data/NNRat/Lemmas.lean | 1 + .../GroupTheory/GroupAction/DomAct/Basic.lean | 1 + 11 files changed, 164 insertions(+), 115 deletions(-) create mode 100644 Mathlib/Algebra/GroupWithZero/Action/End.lean create mode 100644 Mathlib/Algebra/GroupWithZero/Action/Faithful.lean diff --git a/Mathlib.lean b/Mathlib.lean index e839e2377857a..23c1f16aa849d 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -319,6 +319,8 @@ import Mathlib.Algebra.Group.ZeroOne import Mathlib.Algebra.GroupPower.IterateHom import Mathlib.Algebra.GroupWithZero.Action.Basic import Mathlib.Algebra.GroupWithZero.Action.Defs +import Mathlib.Algebra.GroupWithZero.Action.End +import Mathlib.Algebra.GroupWithZero.Action.Faithful import Mathlib.Algebra.GroupWithZero.Action.Opposite import Mathlib.Algebra.GroupWithZero.Action.Pi import Mathlib.Algebra.GroupWithZero.Action.Prod diff --git a/Mathlib/Algebra/Group/Submonoid/DistribMulAction.lean b/Mathlib/Algebra/Group/Submonoid/DistribMulAction.lean index d11f9c6414e1f..15e285f0d6b07 100644 --- a/Mathlib/Algebra/Group/Submonoid/DistribMulAction.lean +++ b/Mathlib/Algebra/Group/Submonoid/DistribMulAction.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ import Mathlib.Algebra.Group.Submonoid.Operations -import Mathlib.Algebra.GroupWithZero.Action.Defs +import Mathlib.Algebra.GroupWithZero.Action.End /-! # Distributive actions by submonoids diff --git a/Mathlib/Algebra/GroupWithZero/Action/Defs.lean b/Mathlib/Algebra/GroupWithZero/Action/Defs.lean index 0d5f09f7f565a..77eac698ee0d5 100644 --- a/Mathlib/Algebra/GroupWithZero/Action/Defs.lean +++ b/Mathlib/Algebra/GroupWithZero/Action/Defs.lean @@ -3,9 +3,8 @@ Copyright (c) 2018 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Yury Kudryashov -/ -import Mathlib.Algebra.Group.Action.Units -import Mathlib.Algebra.Group.Equiv.Basic -import Mathlib.Algebra.GroupWithZero.Units.Basic +import Mathlib.Algebra.Group.Action.Defs +import Mathlib.Algebra.Group.Hom.Defs /-! # Definitions of group actions @@ -13,25 +12,17 @@ import Mathlib.Algebra.GroupWithZero.Units.Basic This file defines a hierarchy of group action type-classes on top of the previously defined notation classes `SMul` and its additive version `VAdd`: -* `MulAction M α` and its additive version `AddAction G P` are typeclasses used for - actions of multiplicative and additive monoids and groups; they extend notation classes - `SMul` and `VAdd` that are defined in `Algebra.Group.Defs`; +* `SMulZeroClass` is a typeclass for an action that preserves zero +* `DistribSMul M A` is a typeclass for an action on an additive monoid (`AddZeroClass`) that + preserves addition and zero * `DistribMulAction M A` is a typeclass for an action of a multiplicative monoid on an additive monoid such that `a • (b + c) = a • b + a • c` and `a • 0 = 0`. The hierarchy is extended further by `Module`, defined elsewhere. -Also provided are typeclasses for faithful and transitive actions, and typeclasses regarding the -interaction of different group actions, - -* `SMulCommClass M N α` and its additive version `VAddCommClass M N α`; -* `IsScalarTower M N α` and its additive version `VAddAssocClass M N α`; -* `IsCentralScalar M α` and its additive version `IsCentralVAdd M N α`. - ## Notation - `a • b` is used as notation for `SMul.smul a b`. -- `a +ᵥ b` is used as notation for `VAdd.vadd a b`. ## Implementation details @@ -51,42 +42,6 @@ open Function variable {M N A B α β : Type*} -/-- `Monoid.toMulAction` is faithful on nontrivial cancellative monoids with zero. -/ -instance CancelMonoidWithZero.faithfulSMul [CancelMonoidWithZero α] [Nontrivial α] : - FaithfulSMul α α where eq_of_smul_eq_smul h := mul_left_injective₀ one_ne_zero (h 1) - -section GroupWithZero -variable [GroupWithZero α] [MulAction α β] {a : α} - -@[simp] lemma inv_smul_smul₀ (ha : a ≠ 0) (x : β) : a⁻¹ • a • x = x := - inv_smul_smul (Units.mk0 a ha) x - -@[simp] -lemma smul_inv_smul₀ (ha : a ≠ 0) (x : β) : a • a⁻¹ • x = x := smul_inv_smul (Units.mk0 a ha) x - -lemma inv_smul_eq_iff₀ (ha : a ≠ 0) {x y : β} : a⁻¹ • x = y ↔ x = a • y := - inv_smul_eq_iff (g := Units.mk0 a ha) - -lemma eq_inv_smul_iff₀ (ha : a ≠ 0) {x y : β} : x = a⁻¹ • y ↔ a • x = y := - eq_inv_smul_iff (g := Units.mk0 a ha) - -@[simp] -lemma Commute.smul_right_iff₀ [Mul β] [SMulCommClass α β β] [IsScalarTower α β β] {x y : β} - (ha : a ≠ 0) : Commute x (a • y) ↔ Commute x y := Commute.smul_right_iff (g := Units.mk0 a ha) - -@[simp] -lemma Commute.smul_left_iff₀ [Mul β] [SMulCommClass α β β] [IsScalarTower α β β] {x y : β} - (ha : a ≠ 0) : Commute (a • x) y ↔ Commute x y := Commute.smul_left_iff (g := Units.mk0 a ha) - -/-- Right scalar multiplication as an order isomorphism. -/ -@[simps] def Equiv.smulRight (ha : a ≠ 0) : β ≃ β where - toFun b := a • b - invFun b := a⁻¹ • b - left_inv := inv_smul_smul₀ ha - right_inv := smul_inv_smul₀ ha - -end GroupWithZero - /-- Typeclass for scalar multiplication that preserves `0` on the right. -/ class SMulZeroClass (M A : Type*) [Zero A] extends SMul M A where /-- Multiplying `0` by a scalar gives `0` -/ @@ -257,22 +212,8 @@ protected abbrev Function.Surjective.distribMulAction [AddMonoid B] [SMul M B] ( (hf : Surjective f) (smul : ∀ (c : M) (x), f (c • x) = c • f x) : DistribMulAction M B := { hf.distribSMul f smul, hf.mulAction f smul with } -/-- Push forward the action of `R` on `M` along a compatible surjective map `f : R →* S`. - -See also `Function.Surjective.mulActionLeft` and `Function.Surjective.moduleLeft`. --/ -abbrev Function.Surjective.distribMulActionLeft {R S M : Type*} [Monoid R] [AddMonoid M] - [DistribMulAction R M] [Monoid S] [SMul S M] (f : R →* S) (hf : Function.Surjective f) - (hsmul : ∀ (c) (x : M), f c • x = c • x) : DistribMulAction S M := - { hf.distribSMulLeft f hsmul, hf.mulActionLeft f hsmul with } - variable (A) -/-- Compose a `DistribMulAction` with a `MonoidHom`, with action `f r' • m`. -See note [reducible non-instances]. -/ -abbrev DistribMulAction.compHom [Monoid N] (f : N →* M) : DistribMulAction N A := - { DistribSMul.compFun A f, MulAction.compHom A f with } - /-- Each element of the monoid defines an additive monoid homomorphism. -/ @[simps!] def DistribMulAction.toAddMonoidHom (x : M) : A →+ A := @@ -360,13 +301,6 @@ protected abbrev Function.Surjective.mulDistribMulAction [Monoid B] [SMul M B] ( variable (A) -/-- Compose a `MulDistribMulAction` with a `MonoidHom`, with action `f r' • m`. -See note [reducible non-instances]. -/ -abbrev MulDistribMulAction.compHom [Monoid N] (f : N →* M) : MulDistribMulAction N A := - { MulAction.compHom A f with - smul_one := fun x => smul_one (f x), - smul_mul := fun x => smul_mul' (f x) } - /-- Scalar multiplication by `r` as a `MonoidHom`. -/ def MulDistribMulAction.toMonoidHom (r : M) : A →* A where @@ -384,16 +318,6 @@ theorem MulDistribMulAction.toMonoidHom_apply (r : M) (x : A) : @[simp] lemma smul_pow' (r : M) (x : A) (n : ℕ) : r • x ^ n = (r • x) ^ n := (MulDistribMulAction.toMonoidHom _ _).map_pow _ _ -variable (M A) - -/-- Each element of the monoid defines a monoid homomorphism. -/ -@[simps] -def MulDistribMulAction.toMonoidEnd : - M →* Monoid.End A where - toFun := MulDistribMulAction.toMonoidHom A - map_one' := MonoidHom.ext <| one_smul M - map_mul' x y := MonoidHom.ext <| mul_smul x y - end section @@ -409,36 +333,6 @@ theorem smul_div' (r : M) (x y : A) : r • (x / y) = r • x / r • y := end -/-- The tautological action by `AddMonoid.End α` on `α`. - -This generalizes `Function.End.applyMulAction`. -/ -instance AddMonoid.End.applyDistribMulAction [AddMonoid α] : - DistribMulAction (AddMonoid.End α) α where - smul := (· <| ·) - smul_zero := AddMonoidHom.map_zero - smul_add := AddMonoidHom.map_add - one_smul _ := rfl - mul_smul _ _ _ := rfl - -@[simp] -theorem AddMonoid.End.smul_def [AddMonoid α] (f : AddMonoid.End α) (a : α) : f • a = f a := - rfl - -/-- `AddMonoid.End.applyDistribMulAction` is faithful. -/ -instance AddMonoid.End.applyFaithfulSMul [AddMonoid α] : - FaithfulSMul (AddMonoid.End α) α := - ⟨fun {_ _ h} => AddMonoidHom.ext h⟩ - -/-- Each non-zero element of a `GroupWithZero` defines an additive monoid isomorphism of an -`AddMonoid` on which it acts distributively. -This is a stronger version of `DistribMulAction.toAddMonoidHom`. -/ -def DistribMulAction.toAddEquiv₀ {α : Type*} (β : Type*) [GroupWithZero α] [AddMonoid β] - [DistribMulAction α β] (x : α) (hx : x ≠ 0) : β ≃+ β := - { DistribMulAction.toAddMonoidHom β x with - invFun := fun b ↦ x⁻¹ • b - left_inv := fun b ↦ inv_smul_smul₀ hx b - right_inv := fun b ↦ smul_inv_smul₀ hx b } - section Group variable [Group α] [AddMonoid β] [DistribMulAction α β] diff --git a/Mathlib/Algebra/GroupWithZero/Action/End.lean b/Mathlib/Algebra/GroupWithZero/Action/End.lean new file mode 100644 index 0000000000000..e38aaaf10e21d --- /dev/null +++ b/Mathlib/Algebra/GroupWithZero/Action/End.lean @@ -0,0 +1,93 @@ +/- +Copyright (c) 2018 Chris Hughes. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Chris Hughes, Yury Kudryashov +-/ +import Mathlib.Algebra.Group.Action.End +import Mathlib.Algebra.Group.Equiv.Basic +import Mathlib.Algebra.GroupWithZero.Action.Defs +import Mathlib.Algebra.GroupWithZero.Action.Units + +/-! +# Group actions and (endo)morphisms +-/ + +assert_not_exists Equiv.Perm.equivUnitsEnd +assert_not_exists Prod.fst_mul +assert_not_exists Ring + +open Function + +variable {M N A B α β : Type*} + +/-- Push forward the action of `R` on `M` along a compatible surjective map `f : R →* S`. + +See also `Function.Surjective.mulActionLeft` and `Function.Surjective.moduleLeft`. +-/ +abbrev Function.Surjective.distribMulActionLeft {R S M : Type*} [Monoid R] [AddMonoid M] + [DistribMulAction R M] [Monoid S] [SMul S M] (f : R →* S) (hf : Function.Surjective f) + (hsmul : ∀ (c) (x : M), f c • x = c • x) : DistribMulAction S M := + { hf.distribSMulLeft f hsmul, hf.mulActionLeft f hsmul with } + +section AddMonoid + +variable (A) [AddMonoid A] [Monoid M] [DistribMulAction M A] + +/-- Compose a `DistribMulAction` with a `MonoidHom`, with action `f r' • m`. +See note [reducible non-instances]. -/ +abbrev DistribMulAction.compHom [Monoid N] (f : N →* M) : DistribMulAction N A := + { DistribSMul.compFun A f, MulAction.compHom A f with } + +end AddMonoid + +section Monoid + +variable (A) [Monoid A] [Monoid M] [MulDistribMulAction M A] + +/-- Compose a `MulDistribMulAction` with a `MonoidHom`, with action `f r' • m`. +See note [reducible non-instances]. -/ +abbrev MulDistribMulAction.compHom [Monoid N] (f : N →* M) : MulDistribMulAction N A := + { MulAction.compHom A f with + smul_one := fun x => smul_one (f x), + smul_mul := fun x => smul_mul' (f x) } + +end Monoid + +/-- The tautological action by `AddMonoid.End α` on `α`. + +This generalizes `Function.End.applyMulAction`. -/ +instance AddMonoid.End.applyDistribMulAction [AddMonoid α] : + DistribMulAction (AddMonoid.End α) α where + smul := (· <| ·) + smul_zero := AddMonoidHom.map_zero + smul_add := AddMonoidHom.map_add + one_smul _ := rfl + mul_smul _ _ _ := rfl + +@[simp] +theorem AddMonoid.End.smul_def [AddMonoid α] (f : AddMonoid.End α) (a : α) : f • a = f a := + rfl + +/-- `AddMonoid.End.applyDistribMulAction` is faithful. -/ +instance AddMonoid.End.applyFaithfulSMul [AddMonoid α] : + FaithfulSMul (AddMonoid.End α) α := + ⟨fun {_ _ h} => AddMonoidHom.ext h⟩ + +/-- Each non-zero element of a `GroupWithZero` defines an additive monoid isomorphism of an +`AddMonoid` on which it acts distributively. +This is a stronger version of `DistribMulAction.toAddMonoidHom`. -/ +def DistribMulAction.toAddEquiv₀ {α : Type*} (β : Type*) [GroupWithZero α] [AddMonoid β] + [DistribMulAction α β] (x : α) (hx : x ≠ 0) : β ≃+ β := + { DistribMulAction.toAddMonoidHom β x with + invFun := fun b ↦ x⁻¹ • b + left_inv := fun b ↦ inv_smul_smul₀ hx b + right_inv := fun b ↦ smul_inv_smul₀ hx b } + +variable (M A) in +/-- Each element of the monoid defines a monoid homomorphism. -/ +@[simps] +def MulDistribMulAction.toMonoidEnd [Monoid M] [Monoid A] [MulDistribMulAction M A] : + M →* Monoid.End A where + toFun := MulDistribMulAction.toMonoidHom A + map_one' := MonoidHom.ext <| one_smul M + map_mul' x y := MonoidHom.ext <| mul_smul x y diff --git a/Mathlib/Algebra/GroupWithZero/Action/Faithful.lean b/Mathlib/Algebra/GroupWithZero/Action/Faithful.lean new file mode 100644 index 0000000000000..9fb6f8fc8c954 --- /dev/null +++ b/Mathlib/Algebra/GroupWithZero/Action/Faithful.lean @@ -0,0 +1,23 @@ +/- +Copyright (c) 2018 Chris Hughes. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Chris Hughes, Yury Kudryashov +-/ +import Mathlib.Algebra.Group.Action.Faithful +import Mathlib.Algebra.GroupWithZero.NeZero + +/-! +# Faithful actions involving groups with zero +-/ + +assert_not_exists Equiv.Perm.equivUnitsEnd +assert_not_exists Prod.fst_mul +assert_not_exists Ring + +open Function + +variable {M N A B α β : Type*} + +/-- `Monoid.toMulAction` is faithful on nontrivial cancellative monoids with zero. -/ +instance CancelMonoidWithZero.faithfulSMul [CancelMonoidWithZero α] [Nontrivial α] : + FaithfulSMul α α where eq_of_smul_eq_smul h := mul_left_injective₀ one_ne_zero (h 1) diff --git a/Mathlib/Algebra/GroupWithZero/Action/Prod.lean b/Mathlib/Algebra/GroupWithZero/Action/Prod.lean index b825703247594..ff3a88f7e4b13 100644 --- a/Mathlib/Algebra/GroupWithZero/Action/Prod.lean +++ b/Mathlib/Algebra/GroupWithZero/Action/Prod.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Simon Hudon, Patrick Massot, Eric Wieser -/ import Mathlib.Algebra.Group.Action.Prod -import Mathlib.Algebra.GroupWithZero.Action.Defs +import Mathlib.Algebra.GroupWithZero.Action.End /-! # Prod instances for multiplicative actions with zero diff --git a/Mathlib/Algebra/GroupWithZero/Action/Units.lean b/Mathlib/Algebra/GroupWithZero/Action/Units.lean index c517ff3399f1e..5a5df0d53edce 100644 --- a/Mathlib/Algebra/GroupWithZero/Action/Units.lean +++ b/Mathlib/Algebra/GroupWithZero/Action/Units.lean @@ -5,6 +5,7 @@ Authors: Eric Wieser -/ import Mathlib.Algebra.Group.Action.Units import Mathlib.Algebra.GroupWithZero.Action.Defs +import Mathlib.Algebra.GroupWithZero.Units.Basic /-! # Multiplicative actions with zero on and by `Mˣ` @@ -23,7 +24,39 @@ admits a `MulDistribMulAction G Mˣ` structure, again with the obvious definitio * `Algebra.GroupWithZero.Action.Prod` -/ -variable {G M α : Type*} +variable {G M α β : Type*} + +section GroupWithZero +variable [GroupWithZero α] [MulAction α β] {a : α} + +@[simp] lemma inv_smul_smul₀ (ha : a ≠ 0) (x : β) : a⁻¹ • a • x = x := + inv_smul_smul (Units.mk0 a ha) x + +@[simp] +lemma smul_inv_smul₀ (ha : a ≠ 0) (x : β) : a • a⁻¹ • x = x := smul_inv_smul (Units.mk0 a ha) x + +lemma inv_smul_eq_iff₀ (ha : a ≠ 0) {x y : β} : a⁻¹ • x = y ↔ x = a • y := + inv_smul_eq_iff (g := Units.mk0 a ha) + +lemma eq_inv_smul_iff₀ (ha : a ≠ 0) {x y : β} : x = a⁻¹ • y ↔ a • x = y := + eq_inv_smul_iff (g := Units.mk0 a ha) + +@[simp] +lemma Commute.smul_right_iff₀ [Mul β] [SMulCommClass α β β] [IsScalarTower α β β] {x y : β} + (ha : a ≠ 0) : Commute x (a • y) ↔ Commute x y := Commute.smul_right_iff (g := Units.mk0 a ha) + +@[simp] +lemma Commute.smul_left_iff₀ [Mul β] [SMulCommClass α β β] [IsScalarTower α β β] {x y : β} + (ha : a ≠ 0) : Commute (a • x) y ↔ Commute x y := Commute.smul_left_iff (g := Units.mk0 a ha) + +/-- Right scalar multiplication as an order isomorphism. -/ +@[simps] def Equiv.smulRight (ha : a ≠ 0) : β ≃ β where + toFun b := a • b + invFun b := a⁻¹ • b + left_inv := inv_smul_smul₀ ha + right_inv := smul_inv_smul₀ ha + +end GroupWithZero namespace Units diff --git a/Mathlib/Algebra/Module/Defs.lean b/Mathlib/Algebra/Module/Defs.lean index 9b2be9d66ea8b..47b020f71c946 100644 --- a/Mathlib/Algebra/Module/Defs.lean +++ b/Mathlib/Algebra/Module/Defs.lean @@ -3,6 +3,8 @@ Copyright (c) 2015 Nathaniel Thomas. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Nathaniel Thomas, Jeremy Avigad, Johannes Hölzl, Mario Carneiro -/ +import Mathlib.Algebra.GroupWithZero.Action.End +import Mathlib.Algebra.GroupWithZero.Action.Units import Mathlib.Algebra.SMulWithZero import Mathlib.Data.Int.Cast.Lemmas diff --git a/Mathlib/Algebra/Ring/Action/Basic.lean b/Mathlib/Algebra/Ring/Action/Basic.lean index 8466b8daf17c7..d29c1984e45d9 100644 --- a/Mathlib/Algebra/Ring/Action/Basic.lean +++ b/Mathlib/Algebra/Ring/Action/Basic.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Kenny Lau. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau -/ -import Mathlib.Algebra.GroupWithZero.Action.Defs +import Mathlib.Algebra.GroupWithZero.Action.End import Mathlib.Algebra.Ring.Hom.Defs /-! diff --git a/Mathlib/Data/NNRat/Lemmas.lean b/Mathlib/Data/NNRat/Lemmas.lean index 91aadd51015c4..88467e704c1a4 100644 --- a/Mathlib/Data/NNRat/Lemmas.lean +++ b/Mathlib/Data/NNRat/Lemmas.lean @@ -5,6 +5,7 @@ Authors: Yaël Dillies, Bhavik Mehta -/ import Mathlib.Algebra.Field.Rat import Mathlib.Algebra.Group.Indicator +import Mathlib.Algebra.GroupWithZero.Action.End import Mathlib.Algebra.Order.Field.Rat /-! diff --git a/Mathlib/GroupTheory/GroupAction/DomAct/Basic.lean b/Mathlib/GroupTheory/GroupAction/DomAct/Basic.lean index d780cdee10131..9e0ab1536c2d5 100644 --- a/Mathlib/GroupTheory/GroupAction/DomAct/Basic.lean +++ b/Mathlib/GroupTheory/GroupAction/DomAct/Basic.lean @@ -5,6 +5,7 @@ Authors: Yury Kudryashov -/ import Mathlib.Algebra.Group.Opposite import Mathlib.Algebra.Group.Pi.Lemmas +import Mathlib.Algebra.Group.Action.Faithful import Mathlib.Algebra.GroupWithZero.Action.Defs import Mathlib.Algebra.Ring.Defs From f8829a66213bebeab14f58c568eeea9442dd607b Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Fri, 25 Oct 2024 07:22:39 +0000 Subject: [PATCH 402/505] chore: cleanup/split of Data.List.Perm (#18211) Hopefully reduces imports already, and opens the possibility of better splits in `Multiset`. --- Mathlib.lean | 4 +- Mathlib/Algebra/BigOperators/Group/List.lean | 3 +- Mathlib/Data/List/Dedup.lean | 7 + .../Data/List/{Perm.lean => Perm/Basic.lean} | 155 +----------------- Mathlib/Data/List/Perm/Lattice.lean | 107 ++++++++++++ Mathlib/Data/List/Perm/Subperm.lean | 82 +++++++++ Mathlib/Data/List/Permutation.lean | 4 +- Mathlib/Data/List/Prime.lean | 1 - Mathlib/Data/List/Sigma.lean | 3 +- Mathlib/Data/List/Sublists.lean | 2 +- Mathlib/Data/Multiset/Basic.lean | 4 +- Mathlib/Data/Multiset/Dedup.lean | 1 + Mathlib/Data/Nat/Factors.lean | 1 + test/MkIffOfInductive.lean | 2 +- 14 files changed, 216 insertions(+), 160 deletions(-) rename Mathlib/Data/List/{Perm.lean => Perm/Basic.lean} (57%) create mode 100644 Mathlib/Data/List/Perm/Lattice.lean create mode 100644 Mathlib/Data/List/Perm/Subperm.lean diff --git a/Mathlib.lean b/Mathlib.lean index 23c1f16aa849d..d7a779c1cc9bf 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2398,7 +2398,9 @@ import Mathlib.Data.List.NodupEquivFin import Mathlib.Data.List.OfFn import Mathlib.Data.List.Pairwise import Mathlib.Data.List.Palindrome -import Mathlib.Data.List.Perm +import Mathlib.Data.List.Perm.Basic +import Mathlib.Data.List.Perm.Lattice +import Mathlib.Data.List.Perm.Subperm import Mathlib.Data.List.Permutation import Mathlib.Data.List.Pi import Mathlib.Data.List.Prime diff --git a/Mathlib/Algebra/BigOperators/Group/List.lean b/Mathlib/Algebra/BigOperators/Group/List.lean index c61a25c867599..c354261cd7b5d 100644 --- a/Mathlib/Algebra/BigOperators/Group/List.lean +++ b/Mathlib/Algebra/BigOperators/Group/List.lean @@ -8,12 +8,13 @@ import Mathlib.Algebra.Group.Int import Mathlib.Algebra.Group.Nat import Mathlib.Algebra.Group.Opposite import Mathlib.Algebra.Group.Units.Basic -import Mathlib.Data.List.Perm import Mathlib.Data.List.ProdSigma import Mathlib.Data.List.Range import Mathlib.Data.List.Rotate import Mathlib.Data.List.Pairwise import Mathlib.Data.List.Join +import Mathlib.Data.List.Dedup +import Mathlib.Data.List.Perm.Basic /-! # Sums and products from lists diff --git a/Mathlib/Data/List/Dedup.lean b/Mathlib/Data/List/Dedup.lean index 55576850f682a..2c282a581e0bd 100644 --- a/Mathlib/Data/List/Dedup.lean +++ b/Mathlib/Data/List/Dedup.lean @@ -160,4 +160,11 @@ theorem replicate_dedup {x : α} : ∀ {k}, k ≠ 0 → (replicate k x).dedup = theorem count_dedup (l : List α) (a : α) : l.dedup.count a = if a ∈ l then 1 else 0 := by simp_rw [count_eq_of_nodup <| nodup_dedup l, mem_dedup] +theorem Perm.dedup {l₁ l₂ : List α} (p : l₁ ~ l₂) : dedup l₁ ~ dedup l₂ := + perm_iff_count.2 fun a => + if h : a ∈ l₁ then by + simp [h, nodup_dedup, p.subset h] + else by + simp [h, count_eq_zero_of_not_mem, mt p.mem_iff.2] + end List diff --git a/Mathlib/Data/List/Perm.lean b/Mathlib/Data/List/Perm/Basic.lean similarity index 57% rename from Mathlib/Data/List/Perm.lean rename to Mathlib/Data/List/Perm/Basic.lean index efeb8dd594a06..c665732b59f94 100644 --- a/Mathlib/Data/List/Perm.lean +++ b/Mathlib/Data/List/Perm/Basic.lean @@ -3,9 +3,10 @@ Copyright (c) 2015 Microsoft Corporation. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Jeremy Avigad, Mario Carneiro -/ -import Mathlib.Data.List.Dedup -import Mathlib.Data.List.Lattice import Batteries.Data.List.Perm +import Mathlib.Logic.Relation +import Mathlib.Order.RelClasses +import Mathlib.Data.List.Forall2 /-! # List Permutations @@ -97,28 +98,6 @@ theorem rel_perm (hr : BiUnique r) : (Forall₂ r ⇒ Forall₂ r ⇒ (· ↔ · end Rel -section Subperm - -attribute [trans] Subperm.trans - -end Subperm - -lemma subperm_iff : l₁ <+~ l₂ ↔ ∃ l, l ~ l₂ ∧ l₁ <+ l := by - refine ⟨?_, fun ⟨l, h₁, h₂⟩ ↦ h₂.subperm.trans h₁.subperm⟩ - rintro ⟨l, h₁, h₂⟩ - obtain ⟨l', h₂⟩ := h₂.exists_perm_append - exact ⟨l₁ ++ l', (h₂.trans (h₁.append_right _)).symm, (prefix_append _ _).sublist⟩ - -@[simp] lemma subperm_singleton_iff : l <+~ [a] ↔ l = [] ∨ l = [a] := by - constructor - · rw [subperm_iff] - rintro ⟨s, hla, h⟩ - rwa [perm_singleton.mp hla, sublist_singleton] at h - · rintro (rfl | rfl) - exacts [nil_subperm, Subperm.refl _] - -lemma subperm_cons_self : l <+~ a :: l := ⟨l, Perm.refl _, sublist_cons_self _ _⟩ - lemma count_eq_count_filter_add [DecidableEq α] (P : α → Prop) [DecidablePred P] (l : List α) (a : α) : count a l = count a (l.filter P) + count a (l.filter (¬ P ·)) := by @@ -165,70 +144,8 @@ theorem perm_option_toList {o₁ o₂ : Option α} : o₁.toList ~ o₂.toList @[deprecated (since := "2024-10-16")] alias perm_option_to_list := perm_option_toList -alias ⟨subperm.of_cons, subperm.cons⟩ := subperm_cons - --- Porting note: commented out ---attribute [protected] subperm.cons - -theorem cons_subperm_of_mem {a : α} {l₁ l₂ : List α} (d₁ : Nodup l₁) (h₁ : a ∉ l₁) (h₂ : a ∈ l₂) - (s : l₁ <+~ l₂) : a :: l₁ <+~ l₂ := by - rcases s with ⟨l, p, s⟩ - induction s generalizing l₁ with - | slnil => cases h₂ - | @cons r₁ r₂ b s' ih => - simp? at h₂ says simp only [mem_cons] at h₂ - cases' h₂ with e m - · subst b - exact ⟨a :: r₁, p.cons a, s'.cons₂ _⟩ - · rcases ih d₁ h₁ m p with ⟨t, p', s'⟩ - exact ⟨t, p', s'.cons _⟩ - | @cons₂ r₁ r₂ b _ ih => - have bm : b ∈ l₁ := p.subset <| mem_cons_self _ _ - have am : a ∈ r₂ := by - simp only [find?, mem_cons] at h₂ - exact h₂.resolve_left fun e => h₁ <| e.symm ▸ bm - rcases append_of_mem bm with ⟨t₁, t₂, rfl⟩ - have st : t₁ ++ t₂ <+ t₁ ++ b :: t₂ := by simp - rcases ih (d₁.sublist st) (mt (fun x => st.subset x) h₁) am - (Perm.cons_inv <| p.trans perm_middle) with - ⟨t, p', s'⟩ - exact - ⟨b :: t, (p'.cons b).trans <| (swap _ _ _).trans (perm_middle.symm.cons a), s'.cons₂ _⟩ - -protected theorem Nodup.subperm (d : Nodup l₁) (H : l₁ ⊆ l₂) : l₁ <+~ l₂ := - subperm_of_subset d H - -section - -variable [DecidableEq α] - -theorem Perm.bagInter_right {l₁ l₂ : List α} (t : List α) (h : l₁ ~ l₂) : - l₁.bagInter t ~ l₂.bagInter t := by - induction h generalizing t with - | nil => simp - | cons x => by_cases x ∈ t <;> simp [*, Perm.cons] - | swap x y => - by_cases h : x = y - · simp [h] - by_cases xt : x ∈ t <;> by_cases yt : y ∈ t - · simp [xt, yt, mem_erase_of_ne h, mem_erase_of_ne (Ne.symm h), erase_comm, swap] - · simp [xt, yt, mt mem_of_mem_erase, Perm.cons] - · simp [xt, yt, mt mem_of_mem_erase, Perm.cons] - · simp [xt, yt] - | trans _ _ ih_1 ih_2 => exact (ih_1 _).trans (ih_2 _) - -theorem Perm.bagInter_left (l : List α) {t₁ t₂ : List α} (p : t₁ ~ t₂) : - l.bagInter t₁ = l.bagInter t₂ := by - induction' l with a l IH generalizing t₁ t₂ p; · simp - by_cases h : a ∈ t₁ - · simp [h, p.subset h, IH (p.erase _)] - · simp [h, mt p.mem_iff.2 h, IH p] - -theorem Perm.bagInter {l₁ l₂ t₁ t₂ : List α} (hl : l₁ ~ l₂) (ht : t₁ ~ t₂) : - l₁.bagInter t₁ ~ l₂.bagInter t₂ := - ht.bagInter_left l₂ ▸ hl.bagInter_right _ - -theorem perm_replicate_append_replicate {l : List α} {a b : α} {m n : ℕ} (h : a ≠ b) : +theorem perm_replicate_append_replicate + [DecidableEq α] {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 @@ -237,34 +154,6 @@ theorem perm_replicate_append_replicate {l : List α} {a b : α} {m n : ℕ} (h · 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₂ := - perm_iff_count.2 fun a => - if h : a ∈ l₁ then by - simp [h, nodup_dedup, p.subset h] - else by - simp [h, count_eq_zero_of_not_mem, mt p.mem_iff.2] - -theorem Perm.inter_append {l t₁ t₂ : List α} (h : Disjoint t₁ t₂) : - l ∩ (t₁ ++ t₂) ~ l ∩ t₁ ++ l ∩ t₂ := by - induction l with - | nil => simp - | cons x xs l_ih => - by_cases h₁ : x ∈ t₁ - · have h₂ : x ∉ t₂ := h h₁ - simp [*] - by_cases h₂ : x ∈ t₂ - · simp only [*, inter_cons_of_not_mem, false_or, mem_append, inter_cons_of_mem, - not_false_iff] - refine Perm.trans (Perm.cons _ l_ih) ?_ - change [x] ++ xs ∩ t₁ ++ xs ∩ t₂ ~ xs ∩ t₁ ++ ([x] ++ xs ∩ t₂) - rw [← List.append_assoc] - solve_by_elim [Perm.append_right, perm_append_comm] - · simp [*] - -end - - - theorem Perm.bind_left (l : List α) {f g : α → List β} (h : ∀ a ∈ l, f a ~ g a) : l.bind f ~ l.bind g := Perm.join_congr <| by @@ -313,38 +202,4 @@ theorem perm_lookmap (f : α → Option α) {l₁ l₂ : List α} rw [@eq_comm _ y, @eq_comm _ c] apply h d hd c hc -theorem Perm.take_inter [DecidableEq α] {xs ys : List α} (n : ℕ) (h : xs ~ ys) - (h' : ys.Nodup) : xs.take n ~ ys.inter (xs.take n) := by - simp only [List.inter] - exact Perm.trans (show xs.take n ~ xs.filter (xs.take n).elem by - conv_lhs => rw [Nodup.take_eq_filter_mem ((Perm.nodup_iff h).2 h')]) - (Perm.filter _ h) - -theorem Perm.drop_inter [DecidableEq α] {xs ys : List α} (n : ℕ) (h : xs ~ ys) (h' : ys.Nodup) : - xs.drop n ~ ys.inter (xs.drop n) := by - by_cases h'' : n ≤ xs.length - · let n' := xs.length - n - have h₀ : n = xs.length - n' := by rwa [Nat.sub_sub_self] - have h₁ : xs.drop n = (xs.reverse.take n').reverse := by - rw [take_reverse, h₀, reverse_reverse] - rw [h₁] - apply (reverse_perm _).trans - rw [inter_reverse] - apply Perm.take_inter _ _ h' - apply (reverse_perm _).trans; assumption - · have : drop n xs = [] := by - apply eq_nil_of_length_eq_zero - rw [length_drop, Nat.sub_eq_zero_iff_le] - apply le_of_not_ge h'' - simp [this, List.inter] - -theorem Perm.dropSlice_inter [DecidableEq α] {xs ys : List α} (n m : ℕ) (h : xs ~ ys) - (h' : ys.Nodup) : List.dropSlice n m xs ~ ys ∩ List.dropSlice n m xs := by - simp only [dropSlice_eq] - have : n ≤ n + m := Nat.le_add_right _ _ - have h₂ := h.nodup_iff.2 h' - apply Perm.trans _ (Perm.inter_append _).symm - · exact Perm.append (Perm.take_inter _ h h') (Perm.drop_inter _ h h') - · exact disjoint_take_drop h₂ this - end List diff --git a/Mathlib/Data/List/Perm/Lattice.lean b/Mathlib/Data/List/Perm/Lattice.lean new file mode 100644 index 0000000000000..44b3c318a9ff0 --- /dev/null +++ b/Mathlib/Data/List/Perm/Lattice.lean @@ -0,0 +1,107 @@ +/- +Copyright (c) 2015 Microsoft Corporation. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Leonardo de Moura, Jeremy Avigad, Mario Carneiro +-/ +import Mathlib.Data.List.Forall2 +import Mathlib.Data.Set.Pairwise.Basic +import Mathlib.Data.List.Basic +import Mathlib.Data.List.Lattice +import Mathlib.Data.List.Nodup + +/-! +# List Permutations and list lattice operations. + +This file develops theory about the `List.Perm` relation and the lattice structure on lists. +-/ + +-- Make sure we don't import algebra +assert_not_exists Monoid + +open Nat + +namespace List +variable {α β : Type*} {l l₁ l₂ : List α} {a : α} + +open Perm (swap) + +variable [DecidableEq α] + +theorem Perm.bagInter_right {l₁ l₂ : List α} (t : List α) (h : l₁ ~ l₂) : + l₁.bagInter t ~ l₂.bagInter t := by + induction h generalizing t with + | nil => simp + | cons x => by_cases x ∈ t <;> simp [*, Perm.cons] + | swap x y => + by_cases h : x = y + · simp [h] + by_cases xt : x ∈ t <;> by_cases yt : y ∈ t + · simp [xt, yt, mem_erase_of_ne h, mem_erase_of_ne (Ne.symm h), erase_comm, swap] + · simp [xt, yt, mt mem_of_mem_erase, Perm.cons] + · simp [xt, yt, mt mem_of_mem_erase, Perm.cons] + · simp [xt, yt] + | trans _ _ ih_1 ih_2 => exact (ih_1 _).trans (ih_2 _) + +theorem Perm.bagInter_left (l : List α) {t₁ t₂ : List α} (p : t₁ ~ t₂) : + l.bagInter t₁ = l.bagInter t₂ := by + induction' l with a l IH generalizing t₁ t₂ p; · simp + by_cases h : a ∈ t₁ + · simp [h, p.subset h, IH (p.erase _)] + · simp [h, mt p.mem_iff.2 h, IH p] + +theorem Perm.bagInter {l₁ l₂ t₁ t₂ : List α} (hl : l₁ ~ l₂) (ht : t₁ ~ t₂) : + l₁.bagInter t₁ ~ l₂.bagInter t₂ := + ht.bagInter_left l₂ ▸ hl.bagInter_right _ + +theorem Perm.inter_append {l t₁ t₂ : List α} (h : Disjoint t₁ t₂) : + l ∩ (t₁ ++ t₂) ~ l ∩ t₁ ++ l ∩ t₂ := by + induction l with + | nil => simp + | cons x xs l_ih => + by_cases h₁ : x ∈ t₁ + · have h₂ : x ∉ t₂ := h h₁ + simp [*] + by_cases h₂ : x ∈ t₂ + · simp only [*, inter_cons_of_not_mem, false_or, mem_append, inter_cons_of_mem, + not_false_iff] + refine Perm.trans (Perm.cons _ l_ih) ?_ + change [x] ++ xs ∩ t₁ ++ xs ∩ t₂ ~ xs ∩ t₁ ++ ([x] ++ xs ∩ t₂) + rw [← List.append_assoc] + solve_by_elim [Perm.append_right, perm_append_comm] + · simp [*] + +theorem Perm.take_inter {xs ys : List α} (n : ℕ) (h : xs ~ ys) + (h' : ys.Nodup) : xs.take n ~ ys.inter (xs.take n) := by + simp only [List.inter] + exact Perm.trans (show xs.take n ~ xs.filter (xs.take n).elem by + conv_lhs => rw [Nodup.take_eq_filter_mem ((Perm.nodup_iff h).2 h')]) + (Perm.filter _ h) + +theorem Perm.drop_inter {xs ys : List α} (n : ℕ) (h : xs ~ ys) (h' : ys.Nodup) : + xs.drop n ~ ys.inter (xs.drop n) := by + by_cases h'' : n ≤ xs.length + · let n' := xs.length - n + have h₀ : n = xs.length - n' := by rwa [Nat.sub_sub_self] + have h₁ : xs.drop n = (xs.reverse.take n').reverse := by + rw [take_reverse, h₀, reverse_reverse] + rw [h₁] + apply (reverse_perm _).trans + rw [inter_reverse] + apply Perm.take_inter _ _ h' + apply (reverse_perm _).trans; assumption + · have : drop n xs = [] := by + apply eq_nil_of_length_eq_zero + rw [length_drop, Nat.sub_eq_zero_iff_le] + apply le_of_not_ge h'' + simp [this, List.inter] + +theorem Perm.dropSlice_inter {xs ys : List α} (n m : ℕ) (h : xs ~ ys) + (h' : ys.Nodup) : List.dropSlice n m xs ~ ys ∩ List.dropSlice n m xs := by + simp only [dropSlice_eq] + have : n ≤ n + m := Nat.le_add_right _ _ + have h₂ := h.nodup_iff.2 h' + apply Perm.trans _ (Perm.inter_append _).symm + · exact Perm.append (Perm.take_inter _ h h') (Perm.drop_inter _ h h') + · exact disjoint_take_drop h₂ this + +end List diff --git a/Mathlib/Data/List/Perm/Subperm.lean b/Mathlib/Data/List/Perm/Subperm.lean new file mode 100644 index 0000000000000..4093ff882b050 --- /dev/null +++ b/Mathlib/Data/List/Perm/Subperm.lean @@ -0,0 +1,82 @@ +/- +Copyright (c) 2015 Microsoft Corporation. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Leonardo de Moura, Jeremy Avigad, Mario Carneiro +-/ +import Batteries.Data.List.Pairwise +import Batteries.Data.List.Perm +import Mathlib.Data.List.Basic + +/-! +# List Sub-permutations + +This file develops theory about the `List.Subperm` relation. + +## Notation + +The notation `<+~` is used for sub-permutations. +-/ + +open Nat + +namespace List +variable {α β : Type*} {l l₁ l₂ : List α} {a : α} + +open Perm + +section Subperm + +attribute [trans] Subperm.trans + +end Subperm + +lemma subperm_iff : l₁ <+~ l₂ ↔ ∃ l, l ~ l₂ ∧ l₁ <+ l := by + refine ⟨?_, fun ⟨l, h₁, h₂⟩ ↦ h₂.subperm.trans h₁.subperm⟩ + rintro ⟨l, h₁, h₂⟩ + obtain ⟨l', h₂⟩ := h₂.exists_perm_append + exact ⟨l₁ ++ l', (h₂.trans (h₁.append_right _)).symm, (prefix_append _ _).sublist⟩ + +@[simp] lemma subperm_singleton_iff : l <+~ [a] ↔ l = [] ∨ l = [a] := by + constructor + · rw [subperm_iff] + rintro ⟨s, hla, h⟩ + rwa [perm_singleton.mp hla, sublist_singleton] at h + · rintro (rfl | rfl) + exacts [nil_subperm, Subperm.refl _] + +lemma subperm_cons_self : l <+~ a :: l := ⟨l, Perm.refl _, sublist_cons_self _ _⟩ + +alias ⟨subperm.of_cons, subperm.cons⟩ := subperm_cons + +-- Porting note: commented out +--attribute [protected] subperm.cons + +theorem cons_subperm_of_mem {a : α} {l₁ l₂ : List α} (d₁ : Nodup l₁) (h₁ : a ∉ l₁) (h₂ : a ∈ l₂) + (s : l₁ <+~ l₂) : a :: l₁ <+~ l₂ := by + rcases s with ⟨l, p, s⟩ + induction s generalizing l₁ with + | slnil => cases h₂ + | @cons r₁ r₂ b s' ih => + simp? at h₂ says simp only [mem_cons] at h₂ + cases' h₂ with e m + · subst b + exact ⟨a :: r₁, p.cons a, s'.cons₂ _⟩ + · rcases ih d₁ h₁ m p with ⟨t, p', s'⟩ + exact ⟨t, p', s'.cons _⟩ + | @cons₂ r₁ r₂ b _ ih => + have bm : b ∈ l₁ := p.subset <| mem_cons_self _ _ + have am : a ∈ r₂ := by + simp only [find?, mem_cons] at h₂ + exact h₂.resolve_left fun e => h₁ <| e.symm ▸ bm + rcases append_of_mem bm with ⟨t₁, t₂, rfl⟩ + have st : t₁ ++ t₂ <+ t₁ ++ b :: t₂ := by simp + rcases ih (d₁.sublist st) (mt (fun x => st.subset x) h₁) am + (Perm.cons_inv <| p.trans perm_middle) with + ⟨t, p', s'⟩ + exact + ⟨b :: t, (p'.cons b).trans <| (swap _ _ _).trans (perm_middle.symm.cons a), s'.cons₂ _⟩ + +protected theorem Nodup.subperm (d : Nodup l₁) (H : l₁ ⊆ l₂) : l₁ <+~ l₂ := + subperm_of_subset d H + +end List diff --git a/Mathlib/Data/List/Permutation.lean b/Mathlib/Data/List/Permutation.lean index 5bcb74097b405..2798655fad89c 100644 --- a/Mathlib/Data/List/Permutation.lean +++ b/Mathlib/Data/List/Permutation.lean @@ -6,12 +6,10 @@ Authors: Parikshit Khanna, Jeremy Avigad, Leonardo de Moura, Floris van Doorn, M import Mathlib.Data.List.Join import Mathlib.Data.Nat.Factorial.Basic import Mathlib.Data.List.Count -import Mathlib.Data.List.Dedup import Mathlib.Data.List.Duplicate import Mathlib.Data.List.InsertIdx -import Mathlib.Data.List.Lattice -import Mathlib.Data.List.Perm import Batteries.Data.List.Perm +import Mathlib.Data.List.Perm.Basic /-! # Permutations of a list diff --git a/Mathlib/Data/List/Prime.lean b/Mathlib/Data/List/Prime.lean index 45c8f957e4d93..2dce64c46d9df 100644 --- a/Mathlib/Data/List/Prime.lean +++ b/Mathlib/Data/List/Prime.lean @@ -5,7 +5,6 @@ Authors: Johannes Hölzl, Jens Wagemaker, Anne Baanen -/ import Mathlib.Algebra.Associated.Basic import Mathlib.Algebra.BigOperators.Group.List -import Mathlib.Data.List.Perm /-! # Products of lists of prime elements. diff --git a/Mathlib/Data/List/Sigma.lean b/Mathlib/Data/List/Sigma.lean index 06fa67f877f9c..2ca3f245e6cdf 100644 --- a/Mathlib/Data/List/Sigma.lean +++ b/Mathlib/Data/List/Sigma.lean @@ -3,8 +3,9 @@ Copyright (c) 2018 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Sean Leather -/ -import Mathlib.Data.List.Perm import Mathlib.Data.List.Pairwise +import Mathlib.Data.List.Perm.Basic +import Mathlib.Data.List.Nodup /-! # Utilities for lists of sigmas diff --git a/Mathlib/Data/List/Sublists.lean b/Mathlib/Data/List/Sublists.lean index bbe9b5bfedc30..48d8a5c0e48e9 100644 --- a/Mathlib/Data/List/Sublists.lean +++ b/Mathlib/Data/List/Sublists.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ import Mathlib.Data.Nat.Choose.Basic -import Mathlib.Data.List.Perm import Mathlib.Data.List.Range +import Mathlib.Data.List.Perm.Basic /-! # sublists diff --git a/Mathlib/Data/Multiset/Basic.lean b/Mathlib/Data/Multiset/Basic.lean index 667c2e2f5f6f1..56c3a9eefb7ec 100644 --- a/Mathlib/Data/Multiset/Basic.lean +++ b/Mathlib/Data/Multiset/Basic.lean @@ -5,9 +5,11 @@ Authors: Mario Carneiro -/ import Mathlib.Algebra.Group.Nat import Mathlib.Algebra.Order.Sub.Unbundled.Basic -import Mathlib.Data.List.Perm +import Mathlib.Data.List.Perm.Subperm import Mathlib.Data.Set.List import Mathlib.Order.Hom.Basic +import Mathlib.Data.List.Perm.Lattice +import Mathlib.Data.List.Perm.Basic /-! # Multisets diff --git a/Mathlib/Data/Multiset/Dedup.lean b/Mathlib/Data/Multiset/Dedup.lean index 4dee768fea017..1c4737c0e1395 100644 --- a/Mathlib/Data/Multiset/Dedup.lean +++ b/Mathlib/Data/Multiset/Dedup.lean @@ -3,6 +3,7 @@ Copyright (c) 2017 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ +import Mathlib.Data.List.Dedup import Mathlib.Data.Multiset.Nodup /-! diff --git a/Mathlib/Data/Nat/Factors.lean b/Mathlib/Data/Nat/Factors.lean index df65aa467ceea..c3d6f9ad26ecb 100644 --- a/Mathlib/Data/Nat/Factors.lean +++ b/Mathlib/Data/Nat/Factors.lean @@ -8,6 +8,7 @@ import Mathlib.Data.Nat.GCD.Basic import Mathlib.Data.Nat.Prime.Defs import Mathlib.Data.List.Prime import Mathlib.Data.List.Sort +import Mathlib.Data.List.Perm.Subperm /-! # Prime numbers diff --git a/test/MkIffOfInductive.lean b/test/MkIffOfInductive.lean index e2c50aa66a11b..340beef419ff7 100644 --- a/test/MkIffOfInductive.lean +++ b/test/MkIffOfInductive.lean @@ -1,5 +1,5 @@ import Mathlib.Tactic.MkIffOfInductiveProp -import Mathlib.Data.List.Perm +import Mathlib.Data.List.Perm.Lattice mk_iff_of_inductive_prop List.Chain test.chain_iff example {α : Type _} (R : α → α → Prop) (a : α) (al : List α) : From 9d502188e7df20d93aded2b312b28259ff1da308 Mon Sep 17 00:00:00 2001 From: damiano Date: Fri, 25 Oct 2024 07:57:03 +0000 Subject: [PATCH 403/505] chore: remove variables (#18193) #17715 --- Mathlib/Algebra/Module/ZLattice/Covolume.lean | 5 +---- Mathlib/Analysis/Calculus/Deriv/Prod.lean | 7 +------ Mathlib/Analysis/Calculus/FDeriv/Pi.lean | 1 - Mathlib/Analysis/Calculus/FDeriv/Star.lean | 2 +- Mathlib/Analysis/Calculus/Gradient/Basic.lean | 2 +- .../InverseFunctionTheorem/ApproximatesLinearOn.lean | 2 -- Mathlib/Analysis/Calculus/LineDeriv/Basic.lean | 2 +- Mathlib/MeasureTheory/Integral/SetIntegral.lean | 2 +- Mathlib/Order/PrimeIdeal.lean | 2 +- Mathlib/Probability/Distributions/Uniform.lean | 4 ++-- Mathlib/Probability/Kernel/Composition.lean | 6 +++--- Mathlib/Probability/Kernel/Disintegration/CondCDF.lean | 4 ++-- Mathlib/Probability/Kernel/Integral.lean | 2 +- Mathlib/Probability/Kernel/Invariance.lean | 2 +- Mathlib/Probability/Kernel/MeasurableIntegral.lean | 2 +- Mathlib/Probability/Martingale/OptionalSampling.lean | 4 ++-- Mathlib/RingTheory/Flat/Stability.lean | 2 +- Mathlib/RingTheory/HahnSeries/Summable.lean | 2 +- Mathlib/RingTheory/LocalRing/Module.lean | 2 +- Mathlib/RingTheory/Regular/IsSMulRegular.lean | 2 +- Mathlib/Topology/Order/UpperLowerSetTopology.lean | 2 +- 21 files changed, 24 insertions(+), 35 deletions(-) diff --git a/Mathlib/Algebra/Module/ZLattice/Covolume.lean b/Mathlib/Algebra/Module/ZLattice/Covolume.lean index 8b3705ee2aa1e..f5740d29e7505 100644 --- a/Mathlib/Algebra/Module/ZLattice/Covolume.lean +++ b/Mathlib/Algebra/Module/ZLattice/Covolume.lean @@ -33,10 +33,7 @@ open Submodule MeasureTheory Module MeasureTheory Module ZSpan section General -variable (K : Type*) [NormedLinearOrderedField K] [HasSolidNorm K] [FloorRing K] -variable {E : Type*} [NormedAddCommGroup E] [NormedSpace K E] [FiniteDimensional K E] -variable [ProperSpace E] [MeasurableSpace E] -variable (L : Submodule ℤ E) [DiscreteTopology L] [IsZLattice K L] +variable {E : Type*} [NormedAddCommGroup E] [ProperSpace E] [MeasurableSpace E] (L : Submodule ℤ E) /-- The covolume of a `ℤ`-lattice is the volume of some fundamental domain; see `ZLattice.covolume_eq_volume` for the proof that the volume does not depend on the choice of diff --git a/Mathlib/Analysis/Calculus/Deriv/Prod.lean b/Mathlib/Analysis/Calculus/Deriv/Prod.lean index b0640b705f958..dfd2e896128ad 100644 --- a/Mathlib/Analysis/Calculus/Deriv/Prod.lean +++ b/Mathlib/Analysis/Calculus/Deriv/Prod.lean @@ -29,12 +29,7 @@ open Filter Asymptotics Set variable {𝕜 : Type u} [NontriviallyNormedField 𝕜] variable {F : Type v} [NormedAddCommGroup F] [NormedSpace 𝕜 F] -variable {E : Type w} [NormedAddCommGroup E] [NormedSpace 𝕜 E] -variable {f f₀ f₁ g : 𝕜 → F} -variable {f' f₀' f₁' g' : F} -variable {x : 𝕜} -variable {s t : Set 𝕜} -variable {L L₁ L₂ : Filter 𝕜} +variable {f₁ : 𝕜 → F} {f₁' : F} {x : 𝕜} {s : Set 𝕜} {L : Filter 𝕜} section CartesianProduct diff --git a/Mathlib/Analysis/Calculus/FDeriv/Pi.lean b/Mathlib/Analysis/Calculus/FDeriv/Pi.lean index f985ae505ea08..fdc2b96749939 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Pi.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Pi.lean @@ -11,7 +11,6 @@ import Mathlib.Analysis.Calculus.FDeriv.Add variable {𝕜 ι : Type*} [DecidableEq ι] [Fintype ι] [NontriviallyNormedField 𝕜] variable {E : ι → Type*} [∀ i, NormedAddCommGroup (E i)] [∀ i, NormedSpace 𝕜 (E i)] -variable {F : Type*} [NormedAddCommGroup F] [NormedSpace 𝕜 F] @[fun_prop] theorem hasFDerivAt_update (x : ∀ i, E i) {i : ι} (y : E i) : diff --git a/Mathlib/Analysis/Calculus/FDeriv/Star.lean b/Mathlib/Analysis/Calculus/FDeriv/Star.lean index 48f652a8dd13d..933c9c2c4fdc1 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Star.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Star.lean @@ -27,7 +27,7 @@ variable {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] variable {F : Type*} [NormedAddCommGroup F] [StarAddMonoid F] [NormedSpace 𝕜 F] [StarModule 𝕜 F] [ContinuousStar F] -variable {f : E → F} {f' : E →L[𝕜] F} (e : E →L[𝕜] F) {x : E} {s : Set E} {L : Filter E} +variable {f : E → F} {f' : E →L[𝕜] F} {x : E} {s : Set E} {L : Filter E} @[fun_prop] theorem HasStrictFDerivAt.star (h : HasStrictFDerivAt f f' x) : diff --git a/Mathlib/Analysis/Calculus/Gradient/Basic.lean b/Mathlib/Analysis/Calculus/Gradient/Basic.lean index 558b091cb21ac..5ebf3d941d6a8 100644 --- a/Mathlib/Analysis/Calculus/Gradient/Basic.lean +++ b/Mathlib/Analysis/Calculus/Gradient/Basic.lean @@ -252,7 +252,7 @@ section congr /-! ### Congruence properties of the Gradient -/ -variable {f₀ f₁ : F → 𝕜} {f₀' f₁' : F} {x₀ x₁ : F} {s₀ s₁ t : Set F} {L₀ L₁ : Filter F} +variable {f₀ f₁ : F → 𝕜} {f₀' f₁' : F} {t : Set F} theorem Filter.EventuallyEq.hasGradientAtFilter_iff (h₀ : f₀ =ᶠ[L] f₁) (hx : f₀ x = f₁ x) (h₁ : f₀' = f₁') : HasGradientAtFilter f₀ f₀' x L ↔ HasGradientAtFilter f₁ f₁' x L := diff --git a/Mathlib/Analysis/Calculus/InverseFunctionTheorem/ApproximatesLinearOn.lean b/Mathlib/Analysis/Calculus/InverseFunctionTheorem/ApproximatesLinearOn.lean index 67b3db1e4cd80..b23236aaf0baf 100644 --- a/Mathlib/Analysis/Calculus/InverseFunctionTheorem/ApproximatesLinearOn.lean +++ b/Mathlib/Analysis/Calculus/InverseFunctionTheorem/ApproximatesLinearOn.lean @@ -52,8 +52,6 @@ noncomputable section variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] variable {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] variable {F : Type*} [NormedAddCommGroup F] [NormedSpace 𝕜 F] -variable {G : Type*} [NormedAddCommGroup G] [NormedSpace 𝕜 G] -variable {G' : Type*} [NormedAddCommGroup G'] [NormedSpace 𝕜 G'] variable {ε : ℝ} open Filter Metric Set diff --git a/Mathlib/Analysis/Calculus/LineDeriv/Basic.lean b/Mathlib/Analysis/Calculus/LineDeriv/Basic.lean index 3de1ac70f4558..b675314ee9710 100644 --- a/Mathlib/Analysis/Calculus/LineDeriv/Basic.lean +++ b/Mathlib/Analysis/Calculus/LineDeriv/Basic.lean @@ -475,7 +475,7 @@ section CompRight variable {E : Type*} [AddCommGroup E] [Module 𝕜 E] {E' : Type*} [AddCommGroup E'] [Module 𝕜 E'] - {f : E → F} {f' : F} {x v : E'} {L : E' →ₗ[𝕜] E} + {f : E → F} {f' : F} {x : E'} {L : E' →ₗ[𝕜] E} theorem HasLineDerivAt.of_comp {v : E'} (hf : HasLineDerivAt 𝕜 (f ∘ L) f' x v) : HasLineDerivAt 𝕜 f f' (L x) (L v) := by diff --git a/Mathlib/MeasureTheory/Integral/SetIntegral.lean b/Mathlib/MeasureTheory/Integral/SetIntegral.lean index 6376b28c74713..49b3978cb0252 100644 --- a/Mathlib/MeasureTheory/Integral/SetIntegral.lean +++ b/Mathlib/MeasureTheory/Integral/SetIntegral.lean @@ -66,7 +66,7 @@ variable [MeasurableSpace X] section NormedAddCommGroup variable [NormedAddCommGroup E] [NormedSpace ℝ E] - {f g : X → E} {s t : Set X} {μ ν : Measure X} {l l' : Filter X} + {f g : X → E} {s t : Set X} {μ : Measure X} theorem setIntegral_congr_ae₀ (hs : NullMeasurableSet s μ) (h : ∀ᵐ x ∂μ, x ∈ s → f x = g x) : ∫ x in s, f x ∂μ = ∫ x in s, g x ∂μ := diff --git a/Mathlib/Order/PrimeIdeal.lean b/Mathlib/Order/PrimeIdeal.lean index e8b6364e05f7b..aef2ade7bdfd4 100644 --- a/Mathlib/Order/PrimeIdeal.lean +++ b/Mathlib/Order/PrimeIdeal.lean @@ -101,7 +101,7 @@ end Preorder section SemilatticeInf -variable [SemilatticeInf P] {x y : P} {I : Ideal P} +variable [SemilatticeInf P] {I : Ideal P} theorem IsPrime.mem_or_mem (hI : IsPrime I) {x y : P} : x ⊓ y ∈ I → x ∈ I ∨ y ∈ I := by contrapose! diff --git a/Mathlib/Probability/Distributions/Uniform.lean b/Mathlib/Probability/Distributions/Uniform.lean index 43f54dab8c27d..77f3a13bb4a8c 100644 --- a/Mathlib/Probability/Distributions/Uniform.lean +++ b/Mathlib/Probability/Distributions/Uniform.lean @@ -48,7 +48,7 @@ noncomputable section namespace MeasureTheory -variable {E : Type*} [MeasurableSpace E] {m : Measure E} {μ : Measure E} +variable {E : Type*} [MeasurableSpace E] {μ : Measure E} namespace pdf @@ -206,7 +206,7 @@ end MeasureTheory namespace PMF -variable {α β γ : Type*} +variable {α : Type*} open scoped Classical NNReal ENNReal diff --git a/Mathlib/Probability/Kernel/Composition.lean b/Mathlib/Probability/Kernel/Composition.lean index 9a2db5e2be4bc..5dbebd1c9f25c 100644 --- a/Mathlib/Probability/Kernel/Composition.lean +++ b/Mathlib/Probability/Kernel/Composition.lean @@ -65,7 +65,7 @@ namespace ProbabilityTheory namespace Kernel -variable {α β ι : Type*} {mα : MeasurableSpace α} {mβ : MeasurableSpace β} +variable {α β : Type*} {mα : MeasurableSpace α} {mβ : MeasurableSpace β} section CompositionProduct @@ -335,7 +335,7 @@ end Ae section Restrict -variable {κ : Kernel α β} [IsSFiniteKernel κ] {η : Kernel (α × β) γ} [IsSFiniteKernel η] {a : α} +variable {κ : Kernel α β} [IsSFiniteKernel κ] {η : Kernel (α × β) γ} [IsSFiniteKernel η] theorem compProd_restrict {s : Set β} {t : Set γ} (hs : MeasurableSet s) (ht : MeasurableSet t) : Kernel.restrict κ hs ⊗ₖ Kernel.restrict η ht = Kernel.restrict (κ ⊗ₖ η) (hs.prod ht) := by @@ -759,7 +759,7 @@ def prodMkLeft (γ : Type*) [MeasurableSpace γ] (κ : Kernel α β) : Kernel ( def prodMkRight (γ : Type*) [MeasurableSpace γ] (κ : Kernel α β) : Kernel (α × γ) β := comap κ Prod.fst measurable_fst -variable {γ : Type*} {mγ : MeasurableSpace γ} {f : β → γ} {g : γ → α} +variable {γ : Type*} {mγ : MeasurableSpace γ} @[simp] theorem prodMkLeft_apply (κ : Kernel α β) (ca : γ × α) : prodMkLeft γ κ ca = κ ca.snd := diff --git a/Mathlib/Probability/Kernel/Disintegration/CondCDF.lean b/Mathlib/Probability/Kernel/Disintegration/CondCDF.lean index a5c415461bb5b..c02e122e3cff2 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CondCDF.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CondCDF.lean @@ -42,7 +42,7 @@ open scoped NNReal ENNReal MeasureTheory Topology namespace MeasureTheory.Measure -variable {α β : Type*} {mα : MeasurableSpace α} (ρ : Measure (α × ℝ)) +variable {α : Type*} {mα : MeasurableSpace α} (ρ : Measure (α × ℝ)) /-- Measure on `α` such that for a measurable set `s`, `ρ.IicSnd r s = ρ (s ×ˢ Iic r)`. -/ noncomputable def IicSnd (r : ℝ) : Measure α := @@ -112,7 +112,7 @@ open MeasureTheory namespace ProbabilityTheory -variable {α β ι : Type*} {mα : MeasurableSpace α} +variable {α : Type*} {mα : MeasurableSpace α} attribute [local instance] MeasureTheory.Measure.IsFiniteMeasure.IicSnd diff --git a/Mathlib/Probability/Kernel/Integral.lean b/Mathlib/Probability/Kernel/Integral.lean index f259203acbe03..e97ef25407719 100644 --- a/Mathlib/Probability/Kernel/Integral.lean +++ b/Mathlib/Probability/Kernel/Integral.lean @@ -91,7 +91,7 @@ end Const section Restrict -variable {s t : Set β} +variable {s : Set β} @[simp] theorem integral_restrict (hs : MeasurableSet s) : diff --git a/Mathlib/Probability/Kernel/Invariance.lean b/Mathlib/Probability/Kernel/Invariance.lean index 58d1a2185d16b..e9d40e85a0dd9 100644 --- a/Mathlib/Probability/Kernel/Invariance.lean +++ b/Mathlib/Probability/Kernel/Invariance.lean @@ -30,7 +30,7 @@ open scoped MeasureTheory ENNReal ProbabilityTheory namespace ProbabilityTheory -variable {α β γ : Type*} {mα : MeasurableSpace α} {mβ : MeasurableSpace β} {mγ : MeasurableSpace γ} +variable {α β : Type*} {mα : MeasurableSpace α} {mβ : MeasurableSpace β} namespace Kernel diff --git a/Mathlib/Probability/Kernel/MeasurableIntegral.lean b/Mathlib/Probability/Kernel/MeasurableIntegral.lean index 2d5fff10ca258..63c99fb77ebfa 100644 --- a/Mathlib/Probability/Kernel/MeasurableIntegral.lean +++ b/Mathlib/Probability/Kernel/MeasurableIntegral.lean @@ -231,7 +231,7 @@ alias _root_.Measurable.set_lintegral_kernel := _root_.Measurable.setLIntegral_k end Lintegral -variable {E : Type*} [NormedAddCommGroup E] [IsSFiniteKernel κ] [IsSFiniteKernel η] +variable {E : Type*} [NormedAddCommGroup E] [IsSFiniteKernel κ] theorem measurableSet_kernel_integrable ⦃f : α → β → E⦄ (hf : StronglyMeasurable (uncurry f)) : MeasurableSet {x | Integrable (f x) (κ x)} := by diff --git a/Mathlib/Probability/Martingale/OptionalSampling.lean b/Mathlib/Probability/Martingale/OptionalSampling.lean index 535963b759aba..90b9489aaa39c 100644 --- a/Mathlib/Probability/Martingale/OptionalSampling.lean +++ b/Mathlib/Probability/Martingale/OptionalSampling.lean @@ -143,11 +143,11 @@ and is a measurable space with the Borel σ-algebra. -/ variable {ι : Type*} [LinearOrder ι] [LocallyFiniteOrder ι] [OrderBot ι] [TopologicalSpace ι] [DiscreteTopology ι] [MeasurableSpace ι] [BorelSpace ι] [MeasurableSpace E] [BorelSpace E] - [SecondCountableTopology E] {ℱ : Filtration ι m} {τ σ : Ω → ι} {f : ι → Ω → E} {i n : ι} + [SecondCountableTopology E] {ℱ : Filtration ι m} {τ σ : Ω → ι} {f : ι → Ω → E} {i : ι} theorem condexp_stoppedValue_stopping_time_ae_eq_restrict_le (h : Martingale f ℱ μ) (hτ : IsStoppingTime ℱ τ) (hσ : IsStoppingTime ℱ σ) [SigmaFinite (μ.trim hσ.measurableSpace_le)] - (hτ_le : ∀ x, τ x ≤ n) : + (hτ_le : ∀ x, τ x ≤ i) : μ[stoppedValue f τ|hσ.measurableSpace] =ᵐ[μ.restrict {x : Ω | τ x ≤ σ x}] stoppedValue f τ := by rw [ae_eq_restrict_iff_indicator_ae_eq (hτ.measurableSpace_le _ (hτ.measurableSet_le_stopping_time hσ))] diff --git a/Mathlib/RingTheory/Flat/Stability.lean b/Mathlib/RingTheory/Flat/Stability.lean index 57de113408f23..8b0e1c14b0fa9 100644 --- a/Mathlib/RingTheory/Flat/Stability.lean +++ b/Mathlib/RingTheory/Flat/Stability.lean @@ -153,7 +153,7 @@ section Localization variable {R : Type u} {M Mp : Type*} (Rp : Type v) [CommRing R] [AddCommGroup M] [Module R M] [CommRing Rp] [Algebra R Rp] - [AddCommGroup Mp] [Module R Mp] [Module Rp Mp] [IsScalarTower R Rp Mp] (f : M →ₗ[R] Mp) + [AddCommGroup Mp] [Module R Mp] [Module Rp Mp] [IsScalarTower R Rp Mp] instance localizedModule [Module.Flat R M] (S : Submonoid R) : Module.Flat (Localization S) (LocalizedModule S M) := by diff --git a/Mathlib/RingTheory/HahnSeries/Summable.lean b/Mathlib/RingTheory/HahnSeries/Summable.lean index b4e933391ff50..c3b6bf05ac681 100644 --- a/Mathlib/RingTheory/HahnSeries/Summable.lean +++ b/Mathlib/RingTheory/HahnSeries/Summable.lean @@ -42,7 +42,7 @@ open Pointwise noncomputable section -variable {Γ Γ' R V α β : Type*} +variable {Γ R α β : Type*} namespace HahnSeries diff --git a/Mathlib/RingTheory/LocalRing/Module.lean b/Mathlib/RingTheory/LocalRing/Module.lean index 332bccee7c8f5..d446db6698308 100644 --- a/Mathlib/RingTheory/LocalRing/Module.lean +++ b/Mathlib/RingTheory/LocalRing/Module.lean @@ -28,7 +28,7 @@ This file gathers various results about finite modules over a local ring `(R, `l` is a split injection if and only if `k ⊗ l` is a (split) injection. -/ -variable {R S} [CommRing R] [CommRing S] [Algebra R S] +variable {R} [CommRing R] section diff --git a/Mathlib/RingTheory/Regular/IsSMulRegular.lean b/Mathlib/RingTheory/Regular/IsSMulRegular.lean index efe4d52d7cef9..3619263bfda72 100644 --- a/Mathlib/RingTheory/Regular/IsSMulRegular.lean +++ b/Mathlib/RingTheory/Regular/IsSMulRegular.lean @@ -136,7 +136,7 @@ open Submodule Pointwise variable (M) [CommRing R] [AddCommGroup M] [Module R M] [AddCommGroup M'] [Module R M'] [AddCommGroup M''] [Module R M''] - (I : Ideal R) (N : Submodule R M) (r : R) + (N : Submodule R M) (r : R) variable (R) in lemma biUnion_associatedPrimes_eq_compl_regular [IsNoetherianRing R] : diff --git a/Mathlib/Topology/Order/UpperLowerSetTopology.lean b/Mathlib/Topology/Order/UpperLowerSetTopology.lean index 02681dbd9f86d..c3e5e3a52ad12 100644 --- a/Mathlib/Topology/Order/UpperLowerSetTopology.lean +++ b/Mathlib/Topology/Order/UpperLowerSetTopology.lean @@ -95,7 +95,7 @@ protected def rec {β : WithUpperSet α → Sort*} (h : ∀ a, β (toUpperSet a) instance [Nonempty α] : Nonempty (WithUpperSet α) := ‹Nonempty α› instance [Inhabited α] : Inhabited (WithUpperSet α) := ‹Inhabited α› -variable [Preorder α] [Preorder β] [Preorder γ] +variable [Preorder α] [Preorder β] instance : Preorder (WithUpperSet α) := ‹Preorder α› instance : TopologicalSpace (WithUpperSet α) := upperSet α From d85e7dbb201ed04d86c1f27e373dc6f5fc637092 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Fri, 25 Oct 2024 08:08:23 +0000 Subject: [PATCH 404/505] chore: split Topology.Separation (#18213) Avoids needing GDelta / UniformSpace in many later files. (Identified using #18156) --- Mathlib.lean | 3 +- Mathlib/Analysis/Convex/Radon.lean | 2 +- Mathlib/Dynamics/FixedPoints/Topology.lean | 2 +- .../Topology/Algebra/InfiniteSum/Defs.lean | 2 +- Mathlib/Topology/Algebra/Semigroup.lean | 2 +- .../Topology/Compactification/OnePoint.lean | 1 - Mathlib/Topology/Connected/Separation.lean | 2 +- Mathlib/Topology/CountableSeparatingOn.lean | 2 +- Mathlib/Topology/DenseEmbedding.lean | 2 +- Mathlib/Topology/DiscreteQuotient.lean | 1 - Mathlib/Topology/DiscreteSubset.lean | 2 +- Mathlib/Topology/ExtendFrom.lean | 2 +- Mathlib/Topology/Filter.lean | 2 +- Mathlib/Topology/IndicatorConstPointwise.lean | 2 +- Mathlib/Topology/Instances/Irrational.lean | 1 + Mathlib/Topology/Order/OrderClosed.lean | 2 +- Mathlib/Topology/Order/Priestley.lean | 2 +- Mathlib/Topology/Perfect.lean | 2 +- Mathlib/Topology/QuasiSeparated.lean | 1 - Mathlib/Topology/SeparatedMap.lean | 2 +- .../Basic.lean} | 103 +-------------- Mathlib/Topology/Separation/GDelta.lean | 125 ++++++++++++++++++ Mathlib/Topology/Separation/NotNormal.lean | 1 - Mathlib/Topology/ShrinkingLemma.lean | 2 +- Mathlib/Topology/Sober.lean | 1 - Mathlib/Topology/Specialization.lean | 1 - Mathlib/Topology/Support.lean | 2 +- Mathlib/Topology/UniformSpace/Compact.lean | 1 - Mathlib/Topology/UniformSpace/Separation.lean | 2 +- .../UniformSpace/UniformConvergence.lean | 1 - 30 files changed, 152 insertions(+), 124 deletions(-) rename Mathlib/Topology/{Separation.lean => Separation/Basic.lean} (96%) create mode 100644 Mathlib/Topology/Separation/GDelta.lean diff --git a/Mathlib.lean b/Mathlib.lean index d7a779c1cc9bf..46bd976d0d3e8 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4927,7 +4927,8 @@ import Mathlib.Topology.QuasiSeparated import Mathlib.Topology.RestrictGen import Mathlib.Topology.Semicontinuous import Mathlib.Topology.SeparatedMap -import Mathlib.Topology.Separation +import Mathlib.Topology.Separation.Basic +import Mathlib.Topology.Separation.GDelta import Mathlib.Topology.Separation.NotNormal import Mathlib.Topology.Sequences import Mathlib.Topology.Sets.Closeds diff --git a/Mathlib/Analysis/Convex/Radon.lean b/Mathlib/Analysis/Convex/Radon.lean index 2a65ed2e8986a..21470d3e027c9 100644 --- a/Mathlib/Analysis/Convex/Radon.lean +++ b/Mathlib/Analysis/Convex/Radon.lean @@ -6,7 +6,7 @@ Authors: Vasily Nesterov import Mathlib.Analysis.Convex.Combination import Mathlib.Data.Set.Card import Mathlib.LinearAlgebra.AffineSpace.FiniteDimensional -import Mathlib.Topology.Separation +import Mathlib.Topology.Separation.Basic /-! # Radon's theorem on convex sets diff --git a/Mathlib/Dynamics/FixedPoints/Topology.lean b/Mathlib/Dynamics/FixedPoints/Topology.lean index 092a3114204f4..c657d45d24997 100644 --- a/Mathlib/Dynamics/FixedPoints/Topology.lean +++ b/Mathlib/Dynamics/FixedPoints/Topology.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov, Johannes Hölzl -/ import Mathlib.Dynamics.FixedPoints.Basic -import Mathlib.Topology.Separation +import Mathlib.Topology.Separation.Basic /-! # Topological properties of fixed points diff --git a/Mathlib/Topology/Algebra/InfiniteSum/Defs.lean b/Mathlib/Topology/Algebra/InfiniteSum/Defs.lean index 92ae820ce40a0..2bf685cd201e1 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/Defs.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/Defs.lean @@ -5,7 +5,7 @@ Authors: Johannes Hölzl -/ import Mathlib.Algebra.BigOperators.Finprod import Mathlib.Order.Filter.AtTopBot.BigOperators -import Mathlib.Topology.Separation +import Mathlib.Topology.Separation.Basic /-! # Infinite sum and product over a topological monoid diff --git a/Mathlib/Topology/Algebra/Semigroup.lean b/Mathlib/Topology/Algebra/Semigroup.lean index 9069ed9f766ed..327643f2aa4c3 100644 --- a/Mathlib/Topology/Algebra/Semigroup.lean +++ b/Mathlib/Topology/Algebra/Semigroup.lean @@ -3,8 +3,8 @@ Copyright (c) 2021 David Wärn. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: David Wärn -/ -import Mathlib.Topology.Separation import Mathlib.Algebra.Group.Defs +import Mathlib.Topology.Separation.Basic /-! # Idempotents in topological semigroups diff --git a/Mathlib/Topology/Compactification/OnePoint.lean b/Mathlib/Topology/Compactification/OnePoint.lean index e870661f60d4f..dec6ca2d9e95f 100644 --- a/Mathlib/Topology/Compactification/OnePoint.lean +++ b/Mathlib/Topology/Compactification/OnePoint.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yourong Zang, Yury Kudryashov -/ import Mathlib.Data.Fintype.Option -import Mathlib.Topology.Separation import Mathlib.Topology.Sets.Opens /-! diff --git a/Mathlib/Topology/Connected/Separation.lean b/Mathlib/Topology/Connected/Separation.lean index 8ce21856914c0..f5e3e067c4839 100644 --- a/Mathlib/Topology/Connected/Separation.lean +++ b/Mathlib/Topology/Connected/Separation.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 Dagur Asgeirsson. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Dagur Asgeirsson -/ -import Mathlib.Topology.Separation +import Mathlib.Topology.Separation.Basic /-! # Separation and (dis)connectedness properties of topological spaces. diff --git a/Mathlib/Topology/CountableSeparatingOn.lean b/Mathlib/Topology/CountableSeparatingOn.lean index fd76eddbe41b3..4531e4bdaaa85 100644 --- a/Mathlib/Topology/CountableSeparatingOn.lean +++ b/Mathlib/Topology/CountableSeparatingOn.lean @@ -3,8 +3,8 @@ Copyright (c) 2023 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ -import Mathlib.Topology.Separation import Mathlib.Order.Filter.CountableSeparatingOn +import Mathlib.Topology.Separation.Basic /-! # Countable separating families of sets in topological spaces diff --git a/Mathlib/Topology/DenseEmbedding.lean b/Mathlib/Topology/DenseEmbedding.lean index 210caa3dca529..f42b8f1149a17 100644 --- a/Mathlib/Topology/DenseEmbedding.lean +++ b/Mathlib/Topology/DenseEmbedding.lean @@ -3,8 +3,8 @@ Copyright (c) 2019 Reid Barton. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro, Patrick Massot -/ -import Mathlib.Topology.Separation import Mathlib.Topology.Bases +import Mathlib.Topology.Separation.Basic /-! # Dense embeddings diff --git a/Mathlib/Topology/DiscreteQuotient.lean b/Mathlib/Topology/DiscreteQuotient.lean index 4c20019504bc8..4048f4765b17f 100644 --- a/Mathlib/Topology/DiscreteQuotient.lean +++ b/Mathlib/Topology/DiscreteQuotient.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Calle Sönne, Adam Topaz -/ import Mathlib.Data.Setoid.Partition -import Mathlib.Topology.Separation import Mathlib.Topology.LocallyConstant.Basic /-! diff --git a/Mathlib/Topology/DiscreteSubset.lean b/Mathlib/Topology/DiscreteSubset.lean index 825509afa0b34..70e38c4cc3f54 100644 --- a/Mathlib/Topology/DiscreteSubset.lean +++ b/Mathlib/Topology/DiscreteSubset.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Oliver Nash, Bhavik Mehta, Daniel Weber -/ import Mathlib.Topology.Constructions -import Mathlib.Topology.Separation +import Mathlib.Topology.Separation.Basic /-! # Discrete subsets of topological spaces diff --git a/Mathlib/Topology/ExtendFrom.lean b/Mathlib/Topology/ExtendFrom.lean index 1c043eb2756ea..015142b528778 100644 --- a/Mathlib/Topology/ExtendFrom.lean +++ b/Mathlib/Topology/ExtendFrom.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Anatole Dedecker. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Patrick Massot, Anatole Dedecker -/ -import Mathlib.Topology.Separation +import Mathlib.Topology.Separation.Basic /-! # Extending a function from a subset diff --git a/Mathlib/Topology/Filter.lean b/Mathlib/Topology/Filter.lean index 78b2714743dbe..e2239722200e6 100644 --- a/Mathlib/Topology/Filter.lean +++ b/Mathlib/Topology/Filter.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import Mathlib.Order.Filter.Lift -import Mathlib.Topology.Separation import Mathlib.Order.Interval.Set.Monotone +import Mathlib.Topology.Separation.Basic /-! # Topology on the set of filters on a type diff --git a/Mathlib/Topology/IndicatorConstPointwise.lean b/Mathlib/Topology/IndicatorConstPointwise.lean index c1f6ad354e8ff..deb980b52a7e8 100644 --- a/Mathlib/Topology/IndicatorConstPointwise.lean +++ b/Mathlib/Topology/IndicatorConstPointwise.lean @@ -3,7 +3,7 @@ Copyright (c) 2023 Kalle Kytölä. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kalle Kytölä -/ -import Mathlib.Topology.Separation +import Mathlib.Topology.Separation.Basic /-! # Pointwise convergence of indicator functions diff --git a/Mathlib/Topology/Instances/Irrational.lean b/Mathlib/Topology/Instances/Irrational.lean index 7d4b39145e504..ad7ac6309e9fa 100644 --- a/Mathlib/Topology/Instances/Irrational.lean +++ b/Mathlib/Topology/Instances/Irrational.lean @@ -6,6 +6,7 @@ Authors: Yury Kudryashov import Mathlib.Data.Real.Irrational import Mathlib.Data.Rat.Encodable import Mathlib.Topology.GDelta +import Mathlib.Topology.Separation.GDelta /-! # Topology of irrational numbers diff --git a/Mathlib/Topology/Order/OrderClosed.lean b/Mathlib/Topology/Order/OrderClosed.lean index e2763d592bc24..9cdc99b9b7005 100644 --- a/Mathlib/Topology/Order/OrderClosed.lean +++ b/Mathlib/Topology/Order/OrderClosed.lean @@ -3,7 +3,7 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro, Yury Kudryashov -/ -import Mathlib.Topology.Separation +import Mathlib.Topology.Separation.Basic /-! # Order-closed topologies diff --git a/Mathlib/Topology/Order/Priestley.lean b/Mathlib/Topology/Order/Priestley.lean index ca780c30f0cc8..bdad508392e4a 100644 --- a/Mathlib/Topology/Order/Priestley.lean +++ b/Mathlib/Topology/Order/Priestley.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ import Mathlib.Order.UpperLower.Basic -import Mathlib.Topology.Separation +import Mathlib.Topology.Separation.Basic /-! # Priestley spaces diff --git a/Mathlib/Topology/Perfect.lean b/Mathlib/Topology/Perfect.lean index c390399234943..88e7d579f3dd7 100644 --- a/Mathlib/Topology/Perfect.lean +++ b/Mathlib/Topology/Perfect.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Felix Weilacher -/ -import Mathlib.Topology.Separation +import Mathlib.Topology.Separation.Basic /-! # Perfect Sets diff --git a/Mathlib/Topology/QuasiSeparated.lean b/Mathlib/Topology/QuasiSeparated.lean index 73240b9ea2c7e..f550bb0a8528c 100644 --- a/Mathlib/Topology/QuasiSeparated.lean +++ b/Mathlib/Topology/QuasiSeparated.lean @@ -3,7 +3,6 @@ Copyright (c) 2022 Andrew Yang. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang -/ -import Mathlib.Topology.Separation import Mathlib.Topology.NoetherianSpace /-! diff --git a/Mathlib/Topology/SeparatedMap.lean b/Mathlib/Topology/SeparatedMap.lean index dae896a7a7e4b..992ff27b3e79a 100644 --- a/Mathlib/Topology/SeparatedMap.lean +++ b/Mathlib/Topology/SeparatedMap.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Junyan Xu -/ import Mathlib.Topology.Connected.Basic -import Mathlib.Topology.Separation +import Mathlib.Topology.Separation.Basic /-! # Separated maps and locally injective maps out of a topological space. diff --git a/Mathlib/Topology/Separation.lean b/Mathlib/Topology/Separation/Basic.lean similarity index 96% rename from Mathlib/Topology/Separation.lean rename to Mathlib/Topology/Separation/Basic.lean index 284f7ef7efa3c..cd3e2041e3180 100644 --- a/Mathlib/Topology/Separation.lean +++ b/Mathlib/Topology/Separation/Basic.lean @@ -7,7 +7,6 @@ import Mathlib.Topology.Compactness.Lindelof import Mathlib.Topology.Compactness.SigmaCompact import Mathlib.Topology.Connected.TotallyDisconnected import Mathlib.Topology.Inseparable -import Mathlib.Topology.GDelta /-! # Separation properties of topological spaces. @@ -50,19 +49,13 @@ This file defines the predicate `SeparatedNhds`, and common separation axioms us to conclude that this is equivalent to all subspaces being normal. Such a space is not necessarily Hausdorff or regular, even if it is T₀. * `T5Space`: A T₅ space is a completely normal T₁ space. T₅ implies T₄. -* `PerfectlyNormalSpace`: A perfectly normal space is a normal space such that - closed sets are Gδ. -* `T6Space`: A T₆ space is a Perfectly normal T₁ space. T₆ implies T₅. + +See `Mathlib.Topology.Separation.GDelta` for the definitions of `PerfectlyNormalSpace` and +`T6Space`. Note that `mathlib` adopts the modern convention that `m ≤ n` if and only if `T_m → T_n`, but occasionally the literature swaps definitions for e.g. T₃ and regular. -### TODO - -* Add perfectly normal and T6 spaces. -* Use `hasSeparatingCovers_iff_separatedNhds` to prove that perfectly normal spaces - are completely normal. - ## Main results ### T₀ spaces @@ -124,6 +117,8 @@ If the space is also Lindelöf: -/ +assert_not_exists UniformSpace + open Function Set Filter Topology TopologicalSpace universe u v @@ -900,36 +895,6 @@ instance (priority := 100) ConnectedSpace.neBot_nhdsWithin_compl_of_nontrivial_o contra (compl_union_self _) (Set.nonempty_compl_of_nontrivial _) (singleton_nonempty _) simp [compl_inter_self {x}] at contra -theorem IsGδ.compl_singleton (x : X) [T1Space X] : IsGδ ({x}ᶜ : Set X) := - isOpen_compl_singleton.isGδ - -@[deprecated (since := "2024-02-15")] alias isGδ_compl_singleton := IsGδ.compl_singleton - -theorem Set.Countable.isGδ_compl {s : Set X} [T1Space X] (hs : s.Countable) : IsGδ sᶜ := by - rw [← biUnion_of_singleton s, compl_iUnion₂] - exact .biInter hs fun x _ => .compl_singleton x - -theorem Set.Finite.isGδ_compl {s : Set X} [T1Space X] (hs : s.Finite) : IsGδ sᶜ := - hs.countable.isGδ_compl - -theorem Set.Subsingleton.isGδ_compl {s : Set X} [T1Space X] (hs : s.Subsingleton) : IsGδ sᶜ := - hs.finite.isGδ_compl - -theorem Finset.isGδ_compl [T1Space X] (s : Finset X) : IsGδ (sᶜ : Set X) := - s.finite_toSet.isGδ_compl - -protected theorem IsGδ.singleton [FirstCountableTopology X] [T1Space X] (x : X) : - IsGδ ({x} : Set X) := by - rcases (nhds_basis_opens x).exists_antitone_subbasis with ⟨U, hU, h_basis⟩ - rw [← biInter_basis_nhds h_basis.toHasBasis] - exact .biInter (to_countable _) fun n _ => (hU n).2.isGδ - -@[deprecated (since := "2024-02-15")] alias isGδ_singleton := IsGδ.singleton - -theorem Set.Finite.isGδ [FirstCountableTopology X] {s : Set X} [T1Space X] (hs : s.Finite) : - IsGδ s := - Finite.induction_on hs .empty fun _ _ ↦ .union (.singleton _) - theorem SeparationQuotient.t1Space_iff : T1Space (SeparationQuotient X) ↔ R0Space X := by rw [r0Space_iff, ((t1Space_TFAE (SeparationQuotient X)).out 0 9 :)] constructor @@ -2418,62 +2383,6 @@ instance [CompletelyNormalSpace X] [R0Space X] : T5Space (SeparationQuotient X) end CompletelyNormal -section PerfectlyNormal - -/-- A topological space `X` is a *perfectly normal space* provided it is normal and -closed sets are Gδ. -/ -class PerfectlyNormalSpace (X : Type u) [TopologicalSpace X] extends NormalSpace X : Prop where - closed_gdelta : ∀ ⦃h : Set X⦄, IsClosed h → IsGδ h - -/-- Lemma that allows the easy conclusion that perfectly normal spaces are completely normal. -/ -theorem Disjoint.hasSeparatingCover_closed_gdelta_right {s t : Set X} [NormalSpace X] - (st_dis : Disjoint s t) (t_cl : IsClosed t) (t_gd : IsGδ t) : HasSeparatingCover s t := by - obtain ⟨T, T_open, T_count, T_int⟩ := t_gd - rcases T.eq_empty_or_nonempty with rfl | T_nonempty - · rw [T_int, sInter_empty] at st_dis - rw [(s.disjoint_univ).mp st_dis] - exact t.hasSeparatingCover_empty_left - obtain ⟨g, g_surj⟩ := T_count.exists_surjective T_nonempty - choose g' g'_open clt_sub_g' clg'_sub_g using fun n ↦ by - apply normal_exists_closure_subset t_cl (T_open (g n).1 (g n).2) - rw [T_int] - exact sInter_subset_of_mem (g n).2 - have clg'_int : t = ⋂ i, closure (g' i) := by - apply (subset_iInter fun n ↦ (clt_sub_g' n).trans subset_closure).antisymm - rw [T_int] - refine subset_sInter fun t tinT ↦ ?_ - obtain ⟨n, gn⟩ := g_surj ⟨t, tinT⟩ - refine iInter_subset_of_subset n <| (clg'_sub_g n).trans ?_ - rw [gn] - use fun n ↦ (closure (g' n))ᶜ - constructor - · rw [← compl_iInter, subset_compl_comm, ← clg'_int] - exact st_dis.subset_compl_left - · refine fun n ↦ ⟨isOpen_compl_iff.mpr isClosed_closure, ?_⟩ - simp only [closure_compl, disjoint_compl_left_iff_subset] - rw [← closure_eq_iff_isClosed.mpr t_cl] at clt_sub_g' - exact subset_closure.trans <| (clt_sub_g' n).trans <| (g'_open n).subset_interior_closure - -instance (priority := 100) PerfectlyNormalSpace.toCompletelyNormalSpace - [PerfectlyNormalSpace X] : CompletelyNormalSpace X where - completely_normal _ _ hd₁ hd₂ := separatedNhds_iff_disjoint.mp <| - hasSeparatingCovers_iff_separatedNhds.mp - ⟨(hd₂.hasSeparatingCover_closed_gdelta_right isClosed_closure <| - closed_gdelta isClosed_closure).mono (fun ⦃_⦄ a ↦ a) subset_closure, - ((Disjoint.symm hd₁).hasSeparatingCover_closed_gdelta_right isClosed_closure <| - closed_gdelta isClosed_closure).mono (fun ⦃_⦄ a ↦ a) subset_closure⟩ - -/-- A T₆ space is a perfectly normal T₁ space. -/ -class T6Space (X : Type u) [TopologicalSpace X] extends T1Space X, PerfectlyNormalSpace X : Prop - --- see Note [lower instance priority] -/-- A `T₆` space is a `T₅` space. -/ -instance (priority := 100) T6Space.toT5Space [T6Space X] : T5Space X where - -- follows from type-class inference - -end PerfectlyNormal - - /-- In a compact T₂ space, the connected component of a point equals the intersection of all its clopen neighbourhoods. -/ theorem connectedComponent_eq_iInter_isClopen [T2Space X] [CompactSpace X] (x : X) : @@ -2670,4 +2579,4 @@ instance ConnectedComponents.t2 [T2Space X] [CompactSpace X] : T2Space (Connecte refine ⟨Vᶜ, V, hU.compl.isOpen, hU.isOpen, ?_, hb mem_connectedComponent, disjoint_compl_left⟩ exact fun h => flip Set.Nonempty.ne_empty ha ⟨a, mem_connectedComponent, h⟩ -set_option linter.style.longFile 2800 +set_option linter.style.longFile 2700 diff --git a/Mathlib/Topology/Separation/GDelta.lean b/Mathlib/Topology/Separation/GDelta.lean new file mode 100644 index 0000000000000..3b58de08a18ee --- /dev/null +++ b/Mathlib/Topology/Separation/GDelta.lean @@ -0,0 +1,125 @@ +/- +Copyright (c) 2017 Johannes Hölzl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johannes Hölzl, Mario Carneiro +-/ +import Mathlib.Topology.Compactness.Lindelof +import Mathlib.Topology.Compactness.SigmaCompact +import Mathlib.Topology.Connected.TotallyDisconnected +import Mathlib.Topology.Inseparable +import Mathlib.Topology.GDelta +import Mathlib.Topology.Separation.Basic + +/-! +# Separation properties of topological spaces. + +## Main definitions + +* `PerfectlyNormalSpace`: A perfectly normal space is a normal space such that + closed sets are Gδ. +* `T6Space`: A T₆ space is a Perfectly normal T₁ space. T₆ implies T₅. + +Note that `mathlib` adopts the modern convention that `m ≤ n` if and only if `T_m → T_n`, but +occasionally the literature swaps definitions for e.g. T₃ and regular. + +### TODO + +* Use `hasSeparatingCovers_iff_separatedNhds` to prove that perfectly normal spaces + are completely normal. + +-/ + +open Function Set Filter Topology TopologicalSpace + +universe u v + +variable {X : Type*} {Y : Type*} [TopologicalSpace X] + +section Separation + +theorem IsGδ.compl_singleton (x : X) [T1Space X] : IsGδ ({x}ᶜ : Set X) := + isOpen_compl_singleton.isGδ + +@[deprecated (since := "2024-02-15")] alias isGδ_compl_singleton := IsGδ.compl_singleton + +theorem Set.Countable.isGδ_compl {s : Set X} [T1Space X] (hs : s.Countable) : IsGδ sᶜ := by + rw [← biUnion_of_singleton s, compl_iUnion₂] + exact .biInter hs fun x _ => .compl_singleton x + +theorem Set.Finite.isGδ_compl {s : Set X} [T1Space X] (hs : s.Finite) : IsGδ sᶜ := + hs.countable.isGδ_compl + +theorem Set.Subsingleton.isGδ_compl {s : Set X} [T1Space X] (hs : s.Subsingleton) : IsGδ sᶜ := + hs.finite.isGδ_compl + +theorem Finset.isGδ_compl [T1Space X] (s : Finset X) : IsGδ (sᶜ : Set X) := + s.finite_toSet.isGδ_compl + +protected theorem IsGδ.singleton [FirstCountableTopology X] [T1Space X] (x : X) : + IsGδ ({x} : Set X) := by + rcases (nhds_basis_opens x).exists_antitone_subbasis with ⟨U, hU, h_basis⟩ + rw [← biInter_basis_nhds h_basis.toHasBasis] + exact .biInter (to_countable _) fun n _ => (hU n).2.isGδ + +@[deprecated (since := "2024-02-15")] alias isGδ_singleton := IsGδ.singleton + +theorem Set.Finite.isGδ [FirstCountableTopology X] {s : Set X} [T1Space X] (hs : s.Finite) : + IsGδ s := + Finite.induction_on hs .empty fun _ _ ↦ .union (.singleton _) + + +section PerfectlyNormal + +/-- A topological space `X` is a *perfectly normal space* provided it is normal and +closed sets are Gδ. -/ +class PerfectlyNormalSpace (X : Type u) [TopologicalSpace X] extends NormalSpace X : Prop where + closed_gdelta : ∀ ⦃h : Set X⦄, IsClosed h → IsGδ h + +/-- Lemma that allows the easy conclusion that perfectly normal spaces are completely normal. -/ +theorem Disjoint.hasSeparatingCover_closed_gdelta_right {s t : Set X} [NormalSpace X] + (st_dis : Disjoint s t) (t_cl : IsClosed t) (t_gd : IsGδ t) : HasSeparatingCover s t := by + obtain ⟨T, T_open, T_count, T_int⟩ := t_gd + rcases T.eq_empty_or_nonempty with rfl | T_nonempty + · rw [T_int, sInter_empty] at st_dis + rw [(s.disjoint_univ).mp st_dis] + exact t.hasSeparatingCover_empty_left + obtain ⟨g, g_surj⟩ := T_count.exists_surjective T_nonempty + choose g' g'_open clt_sub_g' clg'_sub_g using fun n ↦ by + apply normal_exists_closure_subset t_cl (T_open (g n).1 (g n).2) + rw [T_int] + exact sInter_subset_of_mem (g n).2 + have clg'_int : t = ⋂ i, closure (g' i) := by + apply (subset_iInter fun n ↦ (clt_sub_g' n).trans subset_closure).antisymm + rw [T_int] + refine subset_sInter fun t tinT ↦ ?_ + obtain ⟨n, gn⟩ := g_surj ⟨t, tinT⟩ + refine iInter_subset_of_subset n <| (clg'_sub_g n).trans ?_ + rw [gn] + use fun n ↦ (closure (g' n))ᶜ + constructor + · rw [← compl_iInter, subset_compl_comm, ← clg'_int] + exact st_dis.subset_compl_left + · refine fun n ↦ ⟨isOpen_compl_iff.mpr isClosed_closure, ?_⟩ + simp only [closure_compl, disjoint_compl_left_iff_subset] + rw [← closure_eq_iff_isClosed.mpr t_cl] at clt_sub_g' + exact subset_closure.trans <| (clt_sub_g' n).trans <| (g'_open n).subset_interior_closure + +instance (priority := 100) PerfectlyNormalSpace.toCompletelyNormalSpace + [PerfectlyNormalSpace X] : CompletelyNormalSpace X where + completely_normal _ _ hd₁ hd₂ := separatedNhds_iff_disjoint.mp <| + hasSeparatingCovers_iff_separatedNhds.mp + ⟨(hd₂.hasSeparatingCover_closed_gdelta_right isClosed_closure <| + closed_gdelta isClosed_closure).mono (fun ⦃_⦄ a ↦ a) subset_closure, + ((Disjoint.symm hd₁).hasSeparatingCover_closed_gdelta_right isClosed_closure <| + closed_gdelta isClosed_closure).mono (fun ⦃_⦄ a ↦ a) subset_closure⟩ + +/-- A T₆ space is a perfectly normal T₁ space. -/ +class T6Space (X : Type u) [TopologicalSpace X] extends T1Space X, PerfectlyNormalSpace X : Prop + +-- see Note [lower instance priority] +/-- A `T₆` space is a `T₅` space. -/ +instance (priority := 100) T6Space.toT5Space [T6Space X] : T5Space X where + -- follows from type-class inference +end PerfectlyNormal + +end Separation diff --git a/Mathlib/Topology/Separation/NotNormal.lean b/Mathlib/Topology/Separation/NotNormal.lean index a022ac86d5359..348f0ea4324cb 100644 --- a/Mathlib/Topology/Separation/NotNormal.lean +++ b/Mathlib/Topology/Separation/NotNormal.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import Mathlib.Data.Real.Cardinality -import Mathlib.Topology.Separation import Mathlib.Topology.TietzeExtension /-! # Not normal topological spaces diff --git a/Mathlib/Topology/ShrinkingLemma.lean b/Mathlib/Topology/ShrinkingLemma.lean index b88d57ecb3233..7cd187f075dfa 100644 --- a/Mathlib/Topology/ShrinkingLemma.lean +++ b/Mathlib/Topology/ShrinkingLemma.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov, Reid Barton -/ -import Mathlib.Topology.Separation +import Mathlib.Topology.Separation.Basic /-! # The shrinking lemma diff --git a/Mathlib/Topology/Sober.lean b/Mathlib/Topology/Sober.lean index 6cf96933116ef..29cbe4476433a 100644 --- a/Mathlib/Topology/Sober.lean +++ b/Mathlib/Topology/Sober.lean @@ -3,7 +3,6 @@ Copyright (c) 2021 Andrew Yang. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang -/ -import Mathlib.Topology.Separation import Mathlib.Topology.Sets.Closeds /-! diff --git a/Mathlib/Topology/Specialization.lean b/Mathlib/Topology/Specialization.lean index 1c7ce45b4bb4b..bd57a351804bd 100644 --- a/Mathlib/Topology/Specialization.lean +++ b/Mathlib/Topology/Specialization.lean @@ -6,7 +6,6 @@ Authors: Yaël Dillies import Mathlib.Order.Category.Preord import Mathlib.Topology.Category.TopCat.Basic import Mathlib.Topology.ContinuousMap.Basic -import Mathlib.Topology.Separation import Mathlib.Topology.Order.UpperLowerSetTopology /-! diff --git a/Mathlib/Topology/Support.lean b/Mathlib/Topology/Support.lean index 30537f87b604c..c533d9dbd075b 100644 --- a/Mathlib/Topology/Support.lean +++ b/Mathlib/Topology/Support.lean @@ -6,7 +6,7 @@ Authors: Floris van Doorn, Patrick Massot import Mathlib.Algebra.GroupWithZero.Indicator import Mathlib.Algebra.Order.Group.Unbundled.Abs import Mathlib.Algebra.Module.Basic -import Mathlib.Topology.Separation +import Mathlib.Topology.Separation.Basic /-! # The topological support of a function diff --git a/Mathlib/Topology/UniformSpace/Compact.lean b/Mathlib/Topology/UniformSpace/Compact.lean index b9bc2156e776b..f8308127388e0 100644 --- a/Mathlib/Topology/UniformSpace/Compact.lean +++ b/Mathlib/Topology/UniformSpace/Compact.lean @@ -5,7 +5,6 @@ Authors: Patrick Massot, Yury Kudryashov -/ import Mathlib.Topology.UniformSpace.UniformConvergence import Mathlib.Topology.UniformSpace.Equicontinuity -import Mathlib.Topology.Separation import Mathlib.Topology.Support /-! diff --git a/Mathlib/Topology/UniformSpace/Separation.lean b/Mathlib/Topology/UniformSpace/Separation.lean index f50de8d5deebd..60e6f6a085671 100644 --- a/Mathlib/Topology/UniformSpace/Separation.lean +++ b/Mathlib/Topology/UniformSpace/Separation.lean @@ -5,7 +5,7 @@ Authors: Johannes Hölzl, Patrick Massot, Yury Kudryashov -/ import Mathlib.Tactic.ApplyFun import Mathlib.Topology.UniformSpace.Basic -import Mathlib.Topology.Separation +import Mathlib.Topology.Separation.Basic /-! # Hausdorff properties of uniform spaces. Separation quotient. diff --git a/Mathlib/Topology/UniformSpace/UniformConvergence.lean b/Mathlib/Topology/UniformSpace/UniformConvergence.lean index 6fe642c53f2e4..aca68d49538ba 100644 --- a/Mathlib/Topology/UniformSpace/UniformConvergence.lean +++ b/Mathlib/Topology/UniformSpace/UniformConvergence.lean @@ -3,7 +3,6 @@ Copyright (c) 2020 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ -import Mathlib.Topology.Separation import Mathlib.Topology.UniformSpace.Basic import Mathlib.Topology.UniformSpace.Cauchy From 1232423bfa0c678e65d76b5970e25b91f58374e6 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Fri, 25 Oct 2024 08:30:34 +0000 Subject: [PATCH 405/505] chore: re-add `@[simp]` attribute lost in the port (#18121) This PR re-adds the `@[simp]` attributes that had a porting note that they could be proved by `@[simp]`, but where that does not appear to be the case currently. Of course, the `simp` set may have changed sufficiently to require un`@[simp]`ing some of these anyway, but we should add a more up to date comment in that case. --- Mathlib/Combinatorics/Young/YoungDiagram.lean | 1 - Mathlib/Geometry/Manifold/MFDeriv/FDeriv.lean | 2 +- Mathlib/GroupTheory/Perm/Sign.lean | 6 +++--- Mathlib/LinearAlgebra/Isomorphisms.lean | 2 +- Mathlib/Logic/Equiv/Defs.lean | 4 ++-- Mathlib/MeasureTheory/Integral/IntervalIntegral.lean | 2 +- Mathlib/MeasureTheory/Measure/NullMeasurable.lean | 1 - Mathlib/RingTheory/AlgebraicIndependent.lean | 12 ++++++++---- Mathlib/RingTheory/MvPowerSeries/Basic.lean | 12 ++++-------- Mathlib/RingTheory/MvPowerSeries/Inverse.lean | 3 +-- Mathlib/RingTheory/PowerSeries/Basic.lean | 6 ++---- Mathlib/RingTheory/PowerSeries/Inverse.lean | 3 +-- Mathlib/SetTheory/Cardinal/ToNat.lean | 3 +-- Mathlib/Topology/Algebra/Module/Basic.lean | 4 ++-- Mathlib/Topology/Category/TopCat/Basic.lean | 4 ++-- Mathlib/Topology/UniformSpace/Equiv.lean | 6 +++--- 16 files changed, 32 insertions(+), 39 deletions(-) diff --git a/Mathlib/Combinatorics/Young/YoungDiagram.lean b/Mathlib/Combinatorics/Young/YoungDiagram.lean index fad6a3949ea28..7bb1431dc4004 100644 --- a/Mathlib/Combinatorics/Young/YoungDiagram.lean +++ b/Mathlib/Combinatorics/Young/YoungDiagram.lean @@ -153,7 +153,6 @@ theorem cells_bot : (⊥ : YoungDiagram).cells = ∅ := rfl -- Porting note: removed `↑`, added `.cells` and changed proof --- @[simp] -- Porting note (#10618): simp can prove this @[norm_cast] theorem coe_bot : (⊥ : YoungDiagram).cells = (∅ : Set (ℕ × ℕ)) := by refine Set.eq_of_subset_of_subset ?_ ?_ diff --git a/Mathlib/Geometry/Manifold/MFDeriv/FDeriv.lean b/Mathlib/Geometry/Manifold/MFDeriv/FDeriv.lean index b0bcd7be1a4a3..40a58337f4b9e 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/FDeriv.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/FDeriv.lean @@ -35,7 +35,7 @@ theorem uniqueMDiffOn_iff_uniqueDiffOn : UniqueMDiffOn 𝓘(𝕜, E) s ↔ Uniqu alias ⟨UniqueMDiffOn.uniqueDiffOn, UniqueDiffOn.uniqueMDiffOn⟩ := uniqueMDiffOn_iff_uniqueDiffOn --- Porting note (#10618): was `@[simp, mfld_simps]` but `simp` can prove it +@[simp, mfld_simps] theorem writtenInExtChartAt_model_space : writtenInExtChartAt 𝓘(𝕜, E) 𝓘(𝕜, E') x f = f := rfl diff --git a/Mathlib/GroupTheory/Perm/Sign.lean b/Mathlib/GroupTheory/Perm/Sign.lean index 98b302cba6c58..41c9a40ec36c5 100644 --- a/Mathlib/GroupTheory/Perm/Sign.lean +++ b/Mathlib/GroupTheory/Perm/Sign.lean @@ -370,7 +370,7 @@ section SignType.sign variable [Fintype α] ---@[simp] Porting note (#10618): simp can prove +@[simp] theorem sign_mul (f g : Perm α) : sign (f * g) = sign f * sign g := MonoidHom.map_mul sign f g @@ -378,7 +378,7 @@ theorem sign_mul (f g : Perm α) : sign (f * g) = sign f * sign g := theorem sign_trans (f g : Perm α) : sign (f.trans g) = sign g * sign f := by rw [← mul_def, sign_mul] ---@[simp] Porting note (#10618): simp can prove +@[simp] theorem sign_one : sign (1 : Perm α) = 1 := MonoidHom.map_one sign @@ -386,7 +386,7 @@ theorem sign_one : sign (1 : Perm α) = 1 := theorem sign_refl : sign (Equiv.refl α) = 1 := MonoidHom.map_one sign ---@[simp] Porting note (#10618): simp can prove +@[simp] theorem sign_inv (f : Perm α) : sign f⁻¹ = sign f := by rw [MonoidHom.map_inv sign f, Int.units_inv_eq_self] diff --git a/Mathlib/LinearAlgebra/Isomorphisms.lean b/Mathlib/LinearAlgebra/Isomorphisms.lean index 0dc3fa092cd9c..3a94028500411 100644 --- a/Mathlib/LinearAlgebra/Isomorphisms.lean +++ b/Mathlib/LinearAlgebra/Isomorphisms.lean @@ -149,7 +149,7 @@ theorem quotientQuotientEquivQuotientAux_mk (x : M ⧸ S) : quotientQuotientEquivQuotientAux S T h (Quotient.mk x) = mapQ S T LinearMap.id h x := liftQ_apply _ _ _ --- @[simp] -- Porting note (#10618): simp can prove this +@[simp] theorem quotientQuotientEquivQuotientAux_mk_mk (x : M) : quotientQuotientEquivQuotientAux S T h (Quotient.mk (Quotient.mk x)) = Quotient.mk x := by simp diff --git a/Mathlib/Logic/Equiv/Defs.lean b/Mathlib/Logic/Equiv/Defs.lean index 0cda134871dce..2650567700550 100644 --- a/Mathlib/Logic/Equiv/Defs.lean +++ b/Mathlib/Logic/Equiv/Defs.lean @@ -599,7 +599,7 @@ theorem psigmaCongrRight_trans {α} {β₁ β₂ β₃ : α → Sort*} theorem psigmaCongrRight_symm {α} {β₁ β₂ : α → Sort*} (F : ∀ a, β₁ a ≃ β₂ a) : (psigmaCongrRight F).symm = psigmaCongrRight fun a => (F a).symm := rfl --- Porting note (#10618): simp can now prove this, so I have removed `@[simp]` +@[simp] theorem psigmaCongrRight_refl {α} {β : α → Sort*} : (psigmaCongrRight fun a => Equiv.refl (β a)) = Equiv.refl (Σ' a, β a) := rfl @@ -620,7 +620,7 @@ theorem sigmaCongrRight_trans {α} {β₁ β₂ β₃ : α → Type*} theorem sigmaCongrRight_symm {α} {β₁ β₂ : α → Type*} (F : ∀ a, β₁ a ≃ β₂ a) : (sigmaCongrRight F).symm = sigmaCongrRight fun a => (F a).symm := rfl --- Porting note (#10618): simp can now prove this, so I have removed `@[simp]` +@[simp] theorem sigmaCongrRight_refl {α} {β : α → Type*} : (sigmaCongrRight fun a => Equiv.refl (β a)) = Equiv.refl (Σ a, β a) := rfl diff --git a/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean b/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean index 6b904ee39dabd..ff624c6f12d1a 100644 --- a/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean +++ b/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean @@ -120,7 +120,7 @@ theorem intervalIntegrable_const_iff {c : E} : IntervalIntegrable (fun _ => c) μ a b ↔ c = 0 ∨ μ (Ι a b) < ∞ := by simp only [intervalIntegrable_iff, integrableOn_const] -@[simp] -- Porting note (#10618): simp can prove this +@[simp] theorem intervalIntegrable_const [IsLocallyFiniteMeasure μ] {c : E} : IntervalIntegrable (fun _ => c) μ a b := intervalIntegrable_const_iff.2 <| Or.inr measure_Ioc_lt_top diff --git a/Mathlib/MeasureTheory/Measure/NullMeasurable.lean b/Mathlib/MeasureTheory/Measure/NullMeasurable.lean index 45d1deff59246..a3090e8782b25 100644 --- a/Mathlib/MeasureTheory/Measure/NullMeasurable.lean +++ b/Mathlib/MeasureTheory/Measure/NullMeasurable.lean @@ -171,7 +171,6 @@ protected theorem disjointed {f : ℕ → Set α} (h : ∀ i, NullMeasurableSet NullMeasurableSet (disjointed f n) μ := MeasurableSet.disjointed h n --- @[simp] -- Porting note (#10618): simp can prove thisrove this protected theorem const (p : Prop) : NullMeasurableSet { _a : α | p } μ := MeasurableSet.const p diff --git a/Mathlib/RingTheory/AlgebraicIndependent.lean b/Mathlib/RingTheory/AlgebraicIndependent.lean index a161c703a9c61..1aa1a3d12eece 100644 --- a/Mathlib/RingTheory/AlgebraicIndependent.lean +++ b/Mathlib/RingTheory/AlgebraicIndependent.lean @@ -363,13 +363,17 @@ theorem AlgebraicIndependent.mvPolynomialOptionEquivPolynomialAdjoin_apply (aeval (fun o : Option ι => o.elim Polynomial.X fun s : ι => Polynomial.C (X s)) y) := rfl ---@[simp] Porting note: removing simp because the linter complains about deterministic timeout +/-- `simp`-normal form of `mvPolynomialOptionEquivPolynomialAdjoin_C` -/ +@[simp] +theorem AlgebraicIndependent.mvPolynomialOptionEquivPolynomialAdjoin_C' + (hx : AlgebraicIndependent R x) (r) : + Polynomial.C (hx.aevalEquiv (C r)) = Polynomial.C (algebraMap _ _ r) := by + simp [aevalEquiv] + theorem AlgebraicIndependent.mvPolynomialOptionEquivPolynomialAdjoin_C (hx : AlgebraicIndependent R x) (r) : hx.mvPolynomialOptionEquivPolynomialAdjoin (C r) = Polynomial.C (algebraMap _ _ r) := by - rw [AlgebraicIndependent.mvPolynomialOptionEquivPolynomialAdjoin_apply, aeval_C, - IsScalarTower.algebraMap_apply R (MvPolynomial ι R), ← Polynomial.C_eq_algebraMap, - Polynomial.map_C, RingHom.coe_coe, AlgEquiv.commutes] + simp theorem AlgebraicIndependent.mvPolynomialOptionEquivPolynomialAdjoin_X_none (hx : AlgebraicIndependent R x) : diff --git a/Mathlib/RingTheory/MvPowerSeries/Basic.lean b/Mathlib/RingTheory/MvPowerSeries/Basic.lean index ab05350f3f87a..a9a882a935dd4 100644 --- a/Mathlib/RingTheory/MvPowerSeries/Basic.lean +++ b/Mathlib/RingTheory/MvPowerSeries/Basic.lean @@ -177,8 +177,7 @@ theorem eq_of_coeff_monomial_ne_zero {m n : σ →₀ ℕ} {a : R} (h : coeff R theorem coeff_comp_monomial (n : σ →₀ ℕ) : (coeff R n).comp (monomial R n) = LinearMap.id := LinearMap.ext <| coeff_monomial_same n --- Porting note (#10618): simp can prove this. --- @[simp] +@[simp] theorem coeff_zero (n : σ →₀ ℕ) : coeff R n (0 : MvPowerSeries σ R) = 0 := rfl @@ -434,13 +433,11 @@ theorem constantCoeff_C (a : R) : constantCoeff σ R (C σ R a) = a := theorem constantCoeff_comp_C : (constantCoeff σ R).comp (C σ R) = RingHom.id R := rfl --- Porting note (#10618): simp can prove this. --- @[simp] +@[simp] theorem constantCoeff_zero : constantCoeff σ R 0 = 0 := rfl --- Porting note (#10618): simp can prove this. --- @[simp] +@[simp] theorem constantCoeff_one : constantCoeff σ R 1 = 1 := rfl @@ -454,8 +451,7 @@ theorem isUnit_constantCoeff (φ : MvPowerSeries σ R) (h : IsUnit φ) : IsUnit (constantCoeff σ R φ) := h.map _ --- Porting note (#10618): simp can prove this. --- @[simp] +@[simp] theorem coeff_smul (f : MvPowerSeries σ R) (n) (a : R) : coeff _ n (a • f) = a * coeff _ n f := rfl diff --git a/Mathlib/RingTheory/MvPowerSeries/Inverse.lean b/Mathlib/RingTheory/MvPowerSeries/Inverse.lean index ce619c30ce099..db252318811bf 100644 --- a/Mathlib/RingTheory/MvPowerSeries/Inverse.lean +++ b/Mathlib/RingTheory/MvPowerSeries/Inverse.lean @@ -228,8 +228,7 @@ theorem inv_eq_zero {φ : MvPowerSeries σ k} : φ⁻¹ = 0 ↔ constantCoeff σ theorem zero_inv : (0 : MvPowerSeries σ k)⁻¹ = 0 := by rw [inv_eq_zero, constantCoeff_zero] --- Porting note (#10618): simp can prove this. --- @[simp] +@[simp] theorem invOfUnit_eq (φ : MvPowerSeries σ k) (h : constantCoeff σ k φ ≠ 0) : invOfUnit φ (Units.mk0 _ h) = φ⁻¹ := rfl diff --git a/Mathlib/RingTheory/PowerSeries/Basic.lean b/Mathlib/RingTheory/PowerSeries/Basic.lean index bae1dcd02e069..65b589c7373de 100644 --- a/Mathlib/RingTheory/PowerSeries/Basic.lean +++ b/Mathlib/RingTheory/PowerSeries/Basic.lean @@ -327,13 +327,11 @@ theorem constantCoeff_C (a : R) : constantCoeff R (C R a) = a := theorem constantCoeff_comp_C : (constantCoeff R).comp (C R) = RingHom.id R := rfl --- Porting note (#10618): simp can prove this. --- @[simp] +@[simp] theorem constantCoeff_zero : constantCoeff R 0 = 0 := rfl --- Porting note (#10618): simp can prove this. --- @[simp] +@[simp] theorem constantCoeff_one : constantCoeff R 1 = 1 := rfl diff --git a/Mathlib/RingTheory/PowerSeries/Inverse.lean b/Mathlib/RingTheory/PowerSeries/Inverse.lean index 4a9202d3df4a7..18c6913c525cd 100644 --- a/Mathlib/RingTheory/PowerSeries/Inverse.lean +++ b/Mathlib/RingTheory/PowerSeries/Inverse.lean @@ -150,8 +150,7 @@ theorem inv_eq_zero {φ : k⟦X⟧} : φ⁻¹ = 0 ↔ constantCoeff k φ = 0 := theorem zero_inv : (0 : k⟦X⟧)⁻¹ = 0 := MvPowerSeries.zero_inv --- Porting note (#10618): simp can prove this. --- @[simp] +@[simp] theorem invOfUnit_eq (φ : k⟦X⟧) (h : constantCoeff k φ ≠ 0) : invOfUnit φ (Units.mk0 _ h) = φ⁻¹ := MvPowerSeries.invOfUnit_eq _ _ diff --git a/Mathlib/SetTheory/Cardinal/ToNat.lean b/Mathlib/SetTheory/Cardinal/ToNat.lean index f37fdaa87cbf3..44953f885edf7 100644 --- a/Mathlib/SetTheory/Cardinal/ToNat.lean +++ b/Mathlib/SetTheory/Cardinal/ToNat.lean @@ -116,8 +116,7 @@ theorem aleph0_toNat : toNat ℵ₀ = 0 := theorem mk_toNat_eq_card [Fintype α] : toNat #α = Fintype.card α := by simp --- porting note (#10618): simp can prove this --- @[simp] +@[simp] theorem zero_toNat : toNat 0 = 0 := map_zero _ theorem one_toNat : toNat 1 = 1 := map_one _ diff --git a/Mathlib/Topology/Algebra/Module/Basic.lean b/Mathlib/Topology/Algebra/Module/Basic.lean index 6293e9dea3c44..fbc19334443c1 100644 --- a/Mathlib/Topology/Algebra/Module/Basic.lean +++ b/Mathlib/Topology/Algebra/Module/Basic.lean @@ -436,7 +436,7 @@ protected theorem map_zero (f : M₁ →SL[σ₁₂] M₂) : f (0 : M₁) = 0 := protected theorem map_add (f : M₁ →SL[σ₁₂] M₂) (x y : M₁) : f (x + y) = f x + f y := map_add f x y --- @[simp] -- Porting note (#10618): simp can prove this +@[simp] protected theorem map_smulₛₗ (f : M₁ →SL[σ₁₂] M₂) (c : R₁) (x : M₁) : f (c • x) = σ₁₂ c • f x := (toLinearMap _).map_smulₛₗ _ _ @@ -1681,7 +1681,7 @@ theorem map_zero (e : M₁ ≃SL[σ₁₂] M₂) : e (0 : M₁) = 0 := theorem map_add (e : M₁ ≃SL[σ₁₂] M₂) (x y : M₁) : e (x + y) = e x + e y := (e : M₁ →SL[σ₁₂] M₂).map_add x y --- @[simp] -- Porting note (#10618): simp can prove this +@[simp] theorem map_smulₛₗ (e : M₁ ≃SL[σ₁₂] M₂) (c : R₁) (x : M₁) : e (c • x) = σ₁₂ c • e x := (e : M₁ →SL[σ₁₂] M₂).map_smulₛₗ c x diff --git a/Mathlib/Topology/Category/TopCat/Basic.lean b/Mathlib/Topology/Category/TopCat/Basic.lean index 57a2905802c66..87760f303fc0e 100644 --- a/Mathlib/Topology/Category/TopCat/Basic.lean +++ b/Mathlib/Topology/Category/TopCat/Basic.lean @@ -53,10 +53,10 @@ instance instFunLike (X Y : TopCat) : FunLike (X ⟶ Y) X Y := instance instContinuousMapClass (X Y : TopCat) : ContinuousMapClass (X ⟶ Y) X Y := inferInstanceAs <| ContinuousMapClass C(X, Y) X Y --- Porting note (#10618): simp can prove this; removed simp +@[simp] theorem id_app (X : TopCat.{u}) (x : ↑X) : (𝟙 X : X ⟶ X) x = x := rfl --- Porting note (#10618): simp can prove this; removed simp +@[simp] theorem comp_app {X Y Z : TopCat.{u}} (f : X ⟶ Y) (g : Y ⟶ Z) (x : X) : (f ≫ g : X → Z) x = g (f x) := rfl diff --git a/Mathlib/Topology/UniformSpace/Equiv.lean b/Mathlib/Topology/UniformSpace/Equiv.lean index 6fb396f7b0004..e9f383c879b3b 100644 --- a/Mathlib/Topology/UniformSpace/Equiv.lean +++ b/Mathlib/Topology/UniformSpace/Equiv.lean @@ -178,7 +178,7 @@ theorem symm_comp_self (h : α ≃ᵤ β) : (h.symm : β → α) ∘ h = id := theorem self_comp_symm (h : α ≃ᵤ β) : (h : α → β) ∘ h.symm = id := funext h.apply_symm_apply --- @[simp] -- Porting note (#10618): `simp` can prove this `simp only [Equiv.range_eq_univ]` +@[simp] theorem range_coe (h : α ≃ᵤ β) : range h = univ := h.surjective.range_eq @@ -188,11 +188,11 @@ theorem image_symm (h : α ≃ᵤ β) : image h.symm = preimage h := theorem preimage_symm (h : α ≃ᵤ β) : preimage h.symm = image h := (funext h.toEquiv.image_eq_preimage).symm --- @[simp] -- Porting note (#10618): `simp` can prove this `simp only [Equiv.image_preimage]` +@[simp] theorem image_preimage (h : α ≃ᵤ β) (s : Set β) : h '' (h ⁻¹' s) = s := h.toEquiv.image_preimage s ---@[simp] -- Porting note (#10618): `simp` can prove this `simp only [Equiv.preimage_image]` +@[simp] theorem preimage_image (h : α ≃ᵤ β) (s : Set α) : h ⁻¹' (h '' s) = s := h.toEquiv.preimage_image s From a3663456d2f9d04002039801a734ff0d65f4b4c7 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Fri, 25 Oct 2024 09:09:02 +0000 Subject: [PATCH 406/505] chore(RingTheory/Kaehler): fix docstring typo (#18196) --- Mathlib/RingTheory/Kaehler/Basic.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/RingTheory/Kaehler/Basic.lean b/Mathlib/RingTheory/Kaehler/Basic.lean index 5492f105c9f9b..d53431a854a9c 100644 --- a/Mathlib/RingTheory/Kaehler/Basic.lean +++ b/Mathlib/RingTheory/Kaehler/Basic.lean @@ -799,7 +799,7 @@ lemma KaehlerDifferential.exact_mapBaseChange_map : end -/-- The map `I → B ⊗[A] B ⊗[A] Ω[A⁄R]` where `I = ker(A → B)`. -/ +/-- The map `I → B ⊗[A] Ω[A⁄R]` where `I = ker(A → B)`. -/ @[simps] noncomputable def KaehlerDifferential.kerToTensor : @@ -811,7 +811,7 @@ def KaehlerDifferential.kerToTensor : algebraMap_eq_smul_one, RingHom.mem_ker.mp x.prop, TensorProduct.zero_tmul, add_zero, RingHom.id_apply] -/-- The map `I/I² → B ⊗[A] B ⊗[A] Ω[A⁄R]` where `I = ker(A → B)`. -/ +/-- The map `I/I² → B ⊗[A] Ω[A⁄R]` where `I = ker(A → B)`. -/ noncomputable def KaehlerDifferential.kerCotangentToTensor : (RingHom.ker (algebraMap A B)).Cotangent →ₗ[A] B ⊗[A] Ω[A⁄R] := From 0418dea8a825428fdeb82bb8d0db769a8c98fa58 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Fri, 25 Oct 2024 09:18:07 +0000 Subject: [PATCH 407/505] chore(Algebra): delete "`simp` can prove this" porting notes (#18056) This PR checks all porting notes of the form "`simp` can prove this" and if this is indeed the case (and the proof appears to make sense), removes the associated porting note. Co-authored-by: Johan Commelin --- Mathlib/Algebra/Algebra/Equiv.lean | 1 - Mathlib/Algebra/Algebra/Hom.lean | 1 - Mathlib/Algebra/Algebra/NonUnitalHom.lean | 3 --- Mathlib/Algebra/Algebra/Operations.lean | 2 -- Mathlib/Algebra/Algebra/Subalgebra/Basic.lean | 1 - Mathlib/Algebra/BigOperators/Finsupp.lean | 4 ---- Mathlib/Algebra/CharP/Defs.lean | 1 - Mathlib/Algebra/CubicDiscriminant.lean | 3 --- Mathlib/Algebra/DirectSum/Ring.lean | 1 - Mathlib/Algebra/Field/Subfield.lean | 2 -- Mathlib/Algebra/GeomSum.lean | 2 -- Mathlib/Algebra/Group/Equiv/Basic.lean | 1 - Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean | 3 --- Mathlib/Algebra/Group/Pointwise/Set/Basic.lean | 1 - Mathlib/Algebra/Group/Subgroup/Basic.lean | 1 - Mathlib/Algebra/Group/Submonoid/Basic.lean | 1 - Mathlib/Algebra/GroupWithZero/Invertible.lean | 1 - Mathlib/Algebra/Homology/HomologicalComplex.lean | 4 ++-- Mathlib/Algebra/Module/Defs.lean | 6 ------ Mathlib/Algebra/Module/LocalizedModule.lean | 2 -- Mathlib/Algebra/Module/Submodule/Range.lean | 1 - Mathlib/Algebra/Module/Torsion.lean | 2 -- Mathlib/Algebra/MonoidAlgebra/Defs.lean | 2 -- Mathlib/Algebra/MonoidAlgebra/Division.lean | 1 - Mathlib/Algebra/MvPolynomial/Expand.lean | 1 - Mathlib/Algebra/MvPolynomial/PDeriv.lean | 1 - Mathlib/Algebra/Order/CauSeq/Completion.lean | 2 -- Mathlib/Algebra/Order/CompleteField.lean | 1 - Mathlib/Algebra/Order/Floor.lean | 4 ---- Mathlib/Algebra/Order/Group/Unbundled/Basic.lean | 2 -- Mathlib/Algebra/Order/GroupWithZero/Canonical.lean | 2 -- Mathlib/Algebra/Order/Monoid/Canonical/Defs.lean | 2 -- Mathlib/Algebra/Order/Monoid/Unbundled/WithTop.lean | 8 -------- Mathlib/Algebra/Polynomial/Basic.lean | 1 - Mathlib/Algebra/Polynomial/Coeff.lean | 4 ++-- Mathlib/Algebra/Polynomial/Degree/Definitions.lean | 3 --- Mathlib/Algebra/Polynomial/Derivative.lean | 5 ----- Mathlib/Algebra/Polynomial/Laurent.lean | 1 - Mathlib/Algebra/Polynomial/Reverse.lean | 2 -- Mathlib/Algebra/Quaternion.lean | 5 ----- Mathlib/Algebra/Ring/BooleanRing.lean | 4 ---- Mathlib/Algebra/Ring/Commute.lean | 4 ---- Mathlib/Algebra/Ring/CompTypeclasses.lean | 6 ------ Mathlib/Algebra/Ring/Equiv.lean | 4 ---- Mathlib/Algebra/Ring/Parity.lean | 10 ---------- Mathlib/Algebra/Ring/Subring/Basic.lean | 4 ---- Mathlib/Algebra/Ring/Subsemiring/Basic.lean | 1 - Mathlib/Algebra/Squarefree/Basic.lean | 1 - Mathlib/Algebra/Symmetrized.lean | 2 -- Mathlib/Algebra/Tropical/Basic.lean | 2 -- 50 files changed, 4 insertions(+), 125 deletions(-) diff --git a/Mathlib/Algebra/Algebra/Equiv.lean b/Mathlib/Algebra/Algebra/Equiv.lean index 6e995f594999b..d29ef71d49492 100644 --- a/Mathlib/Algebra/Algebra/Equiv.lean +++ b/Mathlib/Algebra/Algebra/Equiv.lean @@ -216,7 +216,6 @@ protected theorem map_mul : ∀ x y, e (x * y) = e x * e y := protected theorem map_one : e 1 = 1 := map_one e --- @[simp] -- Porting note (#10618): simp can prove this @[deprecated map_smul (since := "2024-06-20")] protected theorem map_smul (r : R) (x : A₁) : e (r • x) = r • e x := map_smul _ _ _ diff --git a/Mathlib/Algebra/Algebra/Hom.lean b/Mathlib/Algebra/Algebra/Hom.lean index 0f6f35fd279e4..96b78e4f6c6d4 100644 --- a/Mathlib/Algebra/Algebra/Hom.lean +++ b/Mathlib/Algebra/Algebra/Hom.lean @@ -218,7 +218,6 @@ protected theorem map_one : φ 1 = 1 := protected theorem map_pow (x : A) (n : ℕ) : φ (x ^ n) = φ x ^ n := map_pow _ _ _ --- @[simp] -- Porting note (#10618): simp can prove this @[deprecated map_smul (since := "2024-06-26")] protected theorem map_smul (r : R) (x : A) : φ (r • x) = r • φ x := map_smul _ _ _ diff --git a/Mathlib/Algebra/Algebra/NonUnitalHom.lean b/Mathlib/Algebra/Algebra/NonUnitalHom.lean index 678960a93a6bf..157c90299af2a 100644 --- a/Mathlib/Algebra/Algebra/NonUnitalHom.lean +++ b/Mathlib/Algebra/Algebra/NonUnitalHom.lean @@ -245,15 +245,12 @@ theorem coe_mulHom_mk (f : A →ₛₙₐ[φ] B) (h₁ h₂ h₃ h₄) : protected theorem map_smul (f : A →ₛₙₐ[φ] B) (c : R) (x : A) : f (c • x) = (φ c) • f x := map_smulₛₗ _ _ _ --- @[simp] -- Porting note (#10618) : simp can prove this protected theorem map_add (f : A →ₛₙₐ[φ] B) (x y : A) : f (x + y) = f x + f y := map_add _ _ _ --- @[simp] -- Porting note (#10618) : simp can prove this protected theorem map_mul (f : A →ₛₙₐ[φ] B) (x y : A) : f (x * y) = f x * f y := map_mul _ _ _ --- @[simp] -- Porting note (#10618) : simp can prove this protected theorem map_zero (f : A →ₛₙₐ[φ] B) : f 0 = 0 := map_zero _ diff --git a/Mathlib/Algebra/Algebra/Operations.lean b/Mathlib/Algebra/Algebra/Operations.lean index ec2010316a560..ce82654752d9f 100644 --- a/Mathlib/Algebra/Algebra/Operations.lean +++ b/Mathlib/Algebra/Algebra/Operations.lean @@ -184,12 +184,10 @@ theorem mul_bot : M * ⊥ = ⊥ := theorem bot_mul : ⊥ * M = ⊥ := map₂_bot_left _ _ --- @[simp] -- Porting note (#10618): simp can prove this once we have a monoid structure protected theorem one_mul : (1 : Submodule R A) * M = M := by conv_lhs => rw [one_eq_span, ← span_eq M] erw [span_mul_span, one_mul, span_eq] --- @[simp] -- Porting note (#10618): simp can prove this once we have a monoid structure protected theorem mul_one : M * 1 = M := by conv_lhs => rw [one_eq_span, ← span_eq M] erw [span_mul_span, mul_one, span_eq] diff --git a/Mathlib/Algebra/Algebra/Subalgebra/Basic.lean b/Mathlib/Algebra/Algebra/Subalgebra/Basic.lean index 8e6c9f2dc91ce..a34f329d35336 100644 --- a/Mathlib/Algebra/Algebra/Subalgebra/Basic.lean +++ b/Mathlib/Algebra/Algebra/Subalgebra/Basic.lean @@ -46,7 +46,6 @@ instance SubsemiringClass : SubsemiringClass (Subalgebra R A) A where theorem mem_toSubsemiring {S : Subalgebra R A} {x} : x ∈ S.toSubsemiring ↔ x ∈ S := Iff.rfl --- @[simp] -- Porting note (#10618): simp can prove this theorem mem_carrier {s : Subalgebra R A} {x : A} : x ∈ s.carrier ↔ x ∈ s := Iff.rfl diff --git a/Mathlib/Algebra/BigOperators/Finsupp.lean b/Mathlib/Algebra/BigOperators/Finsupp.lean index 7d830fdf78f5a..9b241247cc785 100644 --- a/Mathlib/Algebra/BigOperators/Finsupp.lean +++ b/Mathlib/Algebra/BigOperators/Finsupp.lean @@ -113,8 +113,6 @@ theorem prod_ite_eq' [DecidableEq α] (f : α →₀ M) (a : α) (b : α → M dsimp [Finsupp.prod] rw [f.support.prod_ite_eq'] --- Porting note (#10618): simp can prove this --- @[simp] theorem sum_ite_self_eq' [DecidableEq α] {N : Type*} [AddCommMonoid N] (f : α →₀ N) (a : α) : (f.sum fun x v => ite (x = a) v 0) = f a := by classical @@ -406,8 +404,6 @@ theorem equivFunOnFinite_symm_eq_sum [Fintype α] [AddCommMonoid M] (f : α → ext simp --- Porting note (#10618): simp can prove this --- @[simp] theorem liftAddHom_apply_single [AddCommMonoid M] [AddCommMonoid N] (f : α → M →+ N) (a : α) (b : M) : (liftAddHom (α := α) (M := M) (N := N)) f (single a b) = f a b := sum_single_index (f a).map_zero diff --git a/Mathlib/Algebra/CharP/Defs.lean b/Mathlib/Algebra/CharP/Defs.lean index a80dd345ea9ac..9c41f5a2b3741 100644 --- a/Mathlib/Algebra/CharP/Defs.lean +++ b/Mathlib/Algebra/CharP/Defs.lean @@ -184,7 +184,6 @@ lemma dvd {x : ℕ} (hx : (x : R) = 0) : ringChar R ∣ x := lemma eq_zero [CharZero R] : ringChar R = 0 := eq R 0 --- @[simp] -- Porting note (#10618): simp can prove this lemma Nat.cast_ringChar : (ringChar R : R) = 0 := by rw [ringChar.spec] end ringChar diff --git a/Mathlib/Algebra/CubicDiscriminant.lean b/Mathlib/Algebra/CubicDiscriminant.lean index 7280b1a8b0e49..42e56fed93eff 100644 --- a/Mathlib/Algebra/CubicDiscriminant.lean +++ b/Mathlib/Algebra/CubicDiscriminant.lean @@ -194,7 +194,6 @@ theorem leadingCoeff_of_c_eq_zero (ha : P.a = 0) (hb : P.b = 0) (hc : P.c = 0) : P.toPoly.leadingCoeff = P.d := by rw [of_c_eq_zero ha hb hc, leadingCoeff_C] --- @[simp] -- porting note (#10618): simp can prove this theorem leadingCoeff_of_c_eq_zero' : (toPoly ⟨0, 0, 0, d⟩).leadingCoeff = d := leadingCoeff_of_c_eq_zero rfl rfl rfl @@ -304,7 +303,6 @@ theorem degree_of_d_eq_zero (ha : P.a = 0) (hb : P.b = 0) (hc : P.c = 0) (hd : P P.toPoly.degree = ⊥ := by rw [of_d_eq_zero ha hb hc hd, degree_zero] --- @[simp] -- porting note (#10618): simp can prove this theorem degree_of_d_eq_zero' : (⟨0, 0, 0, 0⟩ : Cubic R).toPoly.degree = ⊥ := degree_of_d_eq_zero rfl rfl rfl rfl @@ -354,7 +352,6 @@ theorem natDegree_of_c_eq_zero (ha : P.a = 0) (hb : P.b = 0) (hc : P.c = 0) : P.toPoly.natDegree = 0 := by rw [of_c_eq_zero ha hb hc, natDegree_C] --- @[simp] -- porting note (#10618): simp can prove this theorem natDegree_of_c_eq_zero' : (toPoly ⟨0, 0, 0, d⟩).natDegree = 0 := natDegree_of_c_eq_zero rfl rfl rfl diff --git a/Mathlib/Algebra/DirectSum/Ring.lean b/Mathlib/Algebra/DirectSum/Ring.lean index b68930289ac05..e30b34a16eff9 100644 --- a/Mathlib/Algebra/DirectSum/Ring.lean +++ b/Mathlib/Algebra/DirectSum/Ring.lean @@ -553,7 +553,6 @@ def toSemiring (f : ∀ i, A i →+ R) (hone : f _ GradedMonoid.GOne.one = 1) simp_rw [of_mul_of, toAddMonoid_of] exact hmul _ _ } --- Porting note (#10618): removed @[simp] as simp can prove this theorem toSemiring_of (f : ∀ i, A i →+ R) (hone hmul) (i : ι) (x : A i) : toSemiring f hone hmul (of _ i x) = f _ x := toAddMonoid_of f i x diff --git a/Mathlib/Algebra/Field/Subfield.lean b/Mathlib/Algebra/Field/Subfield.lean index 30b6731d7f52c..694d6cf7e5504 100644 --- a/Mathlib/Algebra/Field/Subfield.lean +++ b/Mathlib/Algebra/Field/Subfield.lean @@ -176,7 +176,6 @@ instance : SubfieldClass (Subfield K) K where one_mem s := s.one_mem' inv_mem {s} := s.inv_mem' _ --- @[simp] -- Porting note (#10618): simp can prove this (with `coe_toSubring`, which comes later) theorem mem_carrier {s : Subfield K} {x : K} : x ∈ s.carrier ↔ x ∈ s := Iff.rfl @@ -373,7 +372,6 @@ theorem toSubring_subtype_eq_subtype (S : Subfield K) : /-! # Partial order -/ ---@[simp] -- Porting note (#10618): simp can prove this theorem mem_toSubmonoid {s : Subfield K} {x : K} : x ∈ s.toSubmonoid ↔ x ∈ s := Iff.rfl diff --git a/Mathlib/Algebra/GeomSum.lean b/Mathlib/Algebra/GeomSum.lean index 0b34f0c5df0f1..02480271a1392 100644 --- a/Mathlib/Algebra/GeomSum.lean +++ b/Mathlib/Algebra/GeomSum.lean @@ -69,8 +69,6 @@ theorem zero_geom_sum : ∀ {n}, ∑ i ∈ range n, (0 : α) ^ i = if n = 0 then theorem one_geom_sum (n : ℕ) : ∑ i ∈ range n, (1 : α) ^ i = n := by simp --- porting note (#10618): simp can prove this --- @[simp] theorem op_geom_sum (x : α) (n : ℕ) : op (∑ i ∈ range n, x ^ i) = ∑ i ∈ range n, op x ^ i := by simp diff --git a/Mathlib/Algebra/Group/Equiv/Basic.lean b/Mathlib/Algebra/Group/Equiv/Basic.lean index 8d70c2f054093..b11426351ebef 100644 --- a/Mathlib/Algebra/Group/Equiv/Basic.lean +++ b/Mathlib/Algebra/Group/Equiv/Basic.lean @@ -265,7 +265,6 @@ protected theorem injective (e : M ≃* N) : Function.Injective e := protected theorem surjective (e : M ≃* N) : Function.Surjective e := EquivLike.surjective e --- Porting note (#10618): `simp` can prove this @[to_additive] theorem apply_eq_iff_eq (e : M ≃* N) {x y : M} : e x = e y ↔ x = y := e.injective.eq_iff diff --git a/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean b/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean index 5b6bf266bb5de..b23f7252ab9a3 100644 --- a/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean +++ b/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean @@ -708,8 +708,6 @@ theorem div_singleton (a : α) : s / {a} = s.image (· / a) := theorem singleton_div (a : α) : {a} / s = s.image (a / ·) := image₂_singleton_left --- @[to_additive (attr := simp)] --- Porting note (#10618): simp can prove this & the additive version @[to_additive] theorem singleton_div_singleton (a b : α) : ({a} : Finset α) / {b} = {a / b} := image₂_singleton @@ -1251,7 +1249,6 @@ theorem vsub_singleton (b : β) : s -ᵥ ({b} : Finset β) = s.image (· -ᵥ b) theorem singleton_vsub (a : β) : ({a} : Finset β) -ᵥ t = t.image (a -ᵥ ·) := image₂_singleton_left --- @[simp] -- Porting note (#10618): simp can prove this theorem singleton_vsub_singleton (a b : β) : ({a} : Finset β) -ᵥ {b} = {a -ᵥ b} := image₂_singleton diff --git a/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean b/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean index 7dc3e11ce73a6..59cf3b949909d 100644 --- a/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean +++ b/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean @@ -329,7 +329,6 @@ theorem mul_singleton : s * {b} = (· * b) '' s := theorem singleton_mul : {a} * t = (a * ·) '' t := image2_singleton_left --- Porting note (#10618): simp can prove this @[to_additive] theorem singleton_mul_singleton : ({a} : Set α) * {b} = {a * b} := image2_singleton diff --git a/Mathlib/Algebra/Group/Subgroup/Basic.lean b/Mathlib/Algebra/Group/Subgroup/Basic.lean index e53212f0a17ad..9f971900d0349 100644 --- a/Mathlib/Algebra/Group/Subgroup/Basic.lean +++ b/Mathlib/Algebra/Group/Subgroup/Basic.lean @@ -1923,7 +1923,6 @@ theorem normalClosure_eq_iInf : theorem normalClosure_eq_self (H : Subgroup G) [H.Normal] : normalClosure ↑H = H := le_antisymm (normalClosure_le_normal rfl.subset) le_normalClosure --- @[simp] -- Porting note (#10618): simp can prove this theorem normalClosure_idempotent : normalClosure ↑(normalClosure s) = normalClosure s := normalClosure_eq_self _ diff --git a/Mathlib/Algebra/Group/Submonoid/Basic.lean b/Mathlib/Algebra/Group/Submonoid/Basic.lean index 9113ca59629d9..04690c8d30132 100644 --- a/Mathlib/Algebra/Group/Submonoid/Basic.lean +++ b/Mathlib/Algebra/Group/Submonoid/Basic.lean @@ -266,7 +266,6 @@ theorem sup_eq_closure (N N' : Submonoid M) : N ⊔ N' = closure ((N : Set M) theorem closure_iUnion {ι} (s : ι → Set M) : closure (⋃ i, s i) = ⨆ i, closure (s i) := (Submonoid.gi M).gc.l_iSup --- Porting note (#10618): `simp` can now prove this, so we remove the `@[simp]` attribute @[to_additive] theorem closure_singleton_le_iff_mem (m : M) (p : Submonoid M) : closure {m} ≤ p ↔ m ∈ p := by rw [closure_le, singleton_subset_iff, SetLike.mem_coe] diff --git a/Mathlib/Algebra/GroupWithZero/Invertible.lean b/Mathlib/Algebra/GroupWithZero/Invertible.lean index 6895aebeb4888..73b809c2847b2 100644 --- a/Mathlib/Algebra/GroupWithZero/Invertible.lean +++ b/Mathlib/Algebra/GroupWithZero/Invertible.lean @@ -80,7 +80,6 @@ theorem div_self_of_invertible (a : α) [Invertible a] : a / a = 1 := def invertibleDiv (a b : α) [Invertible a] [Invertible b] : Invertible (a / b) := ⟨b / a, by simp [← mul_div_assoc], by simp [← mul_div_assoc]⟩ --- Porting note (#10618): removed `simp` attribute as `simp` can prove it theorem invOf_div (a b : α) [Invertible a] [Invertible b] [Invertible (a / b)] : ⅟ (a / b) = b / a := invOf_eq_right_inv (by simp [← mul_div_assoc]) diff --git a/Mathlib/Algebra/Homology/HomologicalComplex.lean b/Mathlib/Algebra/Homology/HomologicalComplex.lean index 441c75acfbc02..03b17997ea652 100644 --- a/Mathlib/Algebra/Homology/HomologicalComplex.lean +++ b/Mathlib/Algebra/Homology/HomologicalComplex.lean @@ -548,14 +548,14 @@ theorem next_eq (f : Hom C₁ C₂) {i j : ι} (w : c.Rel i j) : obtain rfl := c.next_eq' w simp only [xNextIso, eqToIso_refl, Iso.refl_hom, Iso.refl_inv, comp_id, id_comp] -@[reassoc, elementwise] -- @[simp] -- Porting note (#10618): simp can prove this +@[reassoc, elementwise] theorem comm_from (f : Hom C₁ C₂) (i : ι) : f.f i ≫ C₂.dFrom i = C₁.dFrom i ≫ f.next i := f.comm _ _ attribute [simp 1100] comm_from_assoc attribute [simp] comm_from_apply -@[reassoc, elementwise] -- @[simp] -- Porting note (#10618): simp can prove this +@[reassoc, elementwise] theorem comm_to (f : Hom C₁ C₂) (j : ι) : f.prev j ≫ C₂.dTo j = C₁.dTo j ≫ f.f j := f.comm _ _ diff --git a/Mathlib/Algebra/Module/Defs.lean b/Mathlib/Algebra/Module/Defs.lean index 47b020f71c946..059efa2c28483 100644 --- a/Mathlib/Algebra/Module/Defs.lean +++ b/Mathlib/Algebra/Module/Defs.lean @@ -185,8 +185,6 @@ variable [Ring R] [AddCommGroup M] [Module R M] (r : R) (x : M) theorem neg_smul : -r • x = -(r • x) := eq_neg_of_add_eq_zero_left <| by rw [← add_smul, neg_add_cancel, zero_smul] --- Porting note (#10618): simp can prove this ---@[simp] theorem neg_smul_neg : -r • -x = r • x := by rw [neg_smul, smul_neg, neg_neg] @[simp] @@ -317,13 +315,9 @@ theorem map_natCast_smul [AddCommMonoid M] [AddCommMonoid M₂] {F : Type*} [Fun [Module S M₂] (x : ℕ) (a : M) : f ((x : R) • a) = (x : S) • f a := by simp only [Nat.cast_smul_eq_nsmul, AddMonoidHom.map_nsmul, map_nsmul] --- Porting note (#10618): simp can prove this ---@[simp] theorem Nat.smul_one_eq_cast {R : Type*} [NonAssocSemiring R] (m : ℕ) : m • (1 : R) = ↑m := by rw [nsmul_eq_mul, mul_one] --- Porting note (#10618): simp can prove this ---@[simp] theorem Int.smul_one_eq_cast {R : Type*} [NonAssocRing R] (m : ℤ) : m • (1 : R) = ↑m := by rw [zsmul_eq_mul, mul_one] diff --git a/Mathlib/Algebra/Module/LocalizedModule.lean b/Mathlib/Algebra/Module/LocalizedModule.lean index 231ab81c8c788..936fdf3e75b29 100644 --- a/Mathlib/Algebra/Module/LocalizedModule.lean +++ b/Mathlib/Algebra/Module/LocalizedModule.lean @@ -1003,8 +1003,6 @@ theorem mk'_mul_mk' {M M' : Type*} [Semiring M] [Semiring M'] [Algebra R M] [Alg variable {f} -/-- Porting note (#10618): simp can prove this -@[simp] -/ theorem mk'_eq_iff {m : M} {s : S} {m' : M'} : mk' f m s = m' ↔ f m = s • m' := by rw [← smul_inj f s, Submonoid.smul_def, ← mk'_smul, ← Submonoid.smul_def, mk'_cancel] diff --git a/Mathlib/Algebra/Module/Submodule/Range.lean b/Mathlib/Algebra/Module/Submodule/Range.lean index 44306a514a638..7f70cb40186fa 100644 --- a/Mathlib/Algebra/Module/Submodule/Range.lean +++ b/Mathlib/Algebra/Module/Submodule/Range.lean @@ -266,7 +266,6 @@ theorem map_subtype_le (p' : Submodule R p) : map p.subtype p' ≤ p := by /-- Under the canonical linear map from a submodule `p` to the ambient space `M`, the image of the maximal submodule of `p` is just `p`. -/ --- @[simp] -- Porting note (#10618): simp can prove this theorem map_subtype_top : map p.subtype (⊤ : Submodule R p) = p := by simp @[simp] diff --git a/Mathlib/Algebra/Module/Torsion.lean b/Mathlib/Algebra/Module/Torsion.lean index 5be854af327cb..096deecf0532a 100644 --- a/Mathlib/Algebra/Module/Torsion.lean +++ b/Mathlib/Algebra/Module/Torsion.lean @@ -631,7 +631,6 @@ variable (S : Type*) [CommMonoid S] [DistribMulAction S M] [SMulCommClass S R M] theorem mem_torsion'_iff (x : M) : x ∈ torsion' R M S ↔ ∃ a : S, a • x = 0 := Iff.rfl --- @[simp] Porting note (#10618): simp can prove this theorem mem_torsion_iff (x : M) : x ∈ torsion R M ↔ ∃ a : R⁰, a • x = 0 := Iff.rfl @@ -666,7 +665,6 @@ theorem torsion'_torsion'_eq_top : torsion' R (torsion' R M S) S = ⊤ := /-- The torsion submodule of the torsion submodule (viewed as a module) is the full torsion module. -/ --- @[simp] Porting note (#10618): simp can prove this theorem torsion_torsion_eq_top : torsion R (torsion R M) = ⊤ := torsion'_torsion'_eq_top R⁰ diff --git a/Mathlib/Algebra/MonoidAlgebra/Defs.lean b/Mathlib/Algebra/MonoidAlgebra/Defs.lean index 2b71376c102f0..3766874adfc29 100644 --- a/Mathlib/Algebra/MonoidAlgebra/Defs.lean +++ b/Mathlib/Algebra/MonoidAlgebra/Defs.lean @@ -761,7 +761,6 @@ protected noncomputable def opRingEquiv [Monoid G] : theorem opRingEquiv_single [Monoid G] (r : k) (x : G) : MonoidAlgebra.opRingEquiv (op (single x r)) = single (op x) (op r) := by simp --- @[simp] -- Porting note (#10618): simp can prove this theorem opRingEquiv_symm_single [Monoid G] (r : kᵐᵒᵖ) (x : Gᵐᵒᵖ) : MonoidAlgebra.opRingEquiv.symm (single x r) = op (single x.unop r.unop) := by simp @@ -1421,7 +1420,6 @@ protected noncomputable def opRingEquiv [AddCommMonoid G] : theorem opRingEquiv_single [AddCommMonoid G] (r : k) (x : G) : AddMonoidAlgebra.opRingEquiv (op (single x r)) = single x (op r) := by simp --- @[simp] -- Porting note (#10618): simp can prove this theorem opRingEquiv_symm_single [AddCommMonoid G] (r : kᵐᵒᵖ) (x : Gᵐᵒᵖ) : AddMonoidAlgebra.opRingEquiv.symm (single x r) = op (single x r.unop) := by simp diff --git a/Mathlib/Algebra/MonoidAlgebra/Division.lean b/Mathlib/Algebra/MonoidAlgebra/Division.lean index 42c2c79cdba87..6ed612e48a0a5 100644 --- a/Mathlib/Algebra/MonoidAlgebra/Division.lean +++ b/Mathlib/Algebra/MonoidAlgebra/Division.lean @@ -129,7 +129,6 @@ theorem modOf_apply_of_exists_add (x : k[G]) (g : G) (g' : G) theorem modOf_apply_add_self (x : k[G]) (g : G) (d : G) : (x %ᵒᶠ g) (d + g) = 0 := modOf_apply_of_exists_add _ _ _ ⟨_, add_comm _ _⟩ --- @[simp] -- Porting note (#10618): simp can prove this theorem modOf_apply_self_add (x : k[G]) (g : G) (d : G) : (x %ᵒᶠ g) (g + d) = 0 := modOf_apply_of_exists_add _ _ _ ⟨_, rfl⟩ diff --git a/Mathlib/Algebra/MvPolynomial/Expand.lean b/Mathlib/Algebra/MvPolynomial/Expand.lean index 8b06a208a6ea6..1252dd9b13b66 100644 --- a/Mathlib/Algebra/MvPolynomial/Expand.lean +++ b/Mathlib/Algebra/MvPolynomial/Expand.lean @@ -29,7 +29,6 @@ noncomputable def expand (p : ℕ) : MvPolynomial σ R →ₐ[R] MvPolynomial σ { (eval₂Hom C fun i ↦ X i ^ p : MvPolynomial σ R →+* MvPolynomial σ R) with commutes' := fun _ ↦ eval₂Hom_C _ _ _ } --- @[simp] -- Porting note (#10618): simp can prove this theorem expand_C (p : ℕ) (r : R) : expand p (C r : MvPolynomial σ R) = C r := eval₂Hom_C _ _ _ diff --git a/Mathlib/Algebra/MvPolynomial/PDeriv.lean b/Mathlib/Algebra/MvPolynomial/PDeriv.lean index 2f9c43a1a8601..9c6e2c3c43e7c 100644 --- a/Mathlib/Algebra/MvPolynomial/PDeriv.lean +++ b/Mathlib/Algebra/MvPolynomial/PDeriv.lean @@ -104,7 +104,6 @@ theorem pderiv_pow {i : σ} {f : MvPolynomial σ R} {n : ℕ} : pderiv i (f ^ n) = n * f ^ (n - 1) * pderiv i f := by rw [(pderiv i).leibniz_pow f n, nsmul_eq_mul, smul_eq_mul, mul_assoc] --- @[simp] -- Porting note (#10618): simp can prove this theorem pderiv_C_mul {f : MvPolynomial σ R} {i : σ} : pderiv i (C a * f) = C a * pderiv i f := by rw [C_mul', Derivation.map_smul, C_mul'] diff --git a/Mathlib/Algebra/Order/CauSeq/Completion.lean b/Mathlib/Algebra/Order/CauSeq/Completion.lean index ff8a280c51ac8..501a26686963e 100644 --- a/Mathlib/Algebra/Order/CauSeq/Completion.lean +++ b/Mathlib/Algebra/Order/CauSeq/Completion.lean @@ -198,8 +198,6 @@ noncomputable instance : Inv (Cauchy abv) := rw [mk_eq.2 fg, ← Ig] at If rw [← mul_one (mk (inv f hf)), ← Ig', ← mul_assoc, If, mul_assoc, Ig', mul_one]⟩ --- porting note (#10618): simp can prove this --- @[simp] theorem inv_zero : (0 : (Cauchy abv))⁻¹ = 0 := congr_arg mk <| by rw [dif_pos] <;> [rfl; exact zero_limZero] diff --git a/Mathlib/Algebra/Order/CompleteField.lean b/Mathlib/Algebra/Order/CompleteField.lean index d4a66915b8a75..31a14a28943a9 100644 --- a/Mathlib/Algebra/Order/CompleteField.lean +++ b/Mathlib/Algebra/Order/CompleteField.lean @@ -216,7 +216,6 @@ theorem inducedMap_inducedMap (a : α) : inducedMap β γ (inducedMap α β a) = eq_of_forall_rat_lt_iff_lt fun q => by rw [coe_lt_inducedMap_iff, coe_lt_inducedMap_iff, Iff.comm, coe_lt_inducedMap_iff] ---@[simp] -- Porting note (#10618): simp can prove it theorem inducedMap_inv_self (b : β) : inducedMap γ β (inducedMap β γ b) = b := by rw [inducedMap_inducedMap, inducedMap_self] diff --git a/Mathlib/Algebra/Order/Floor.lean b/Mathlib/Algebra/Order/Floor.lean index da812710841b2..a05e2fffee157 100644 --- a/Mathlib/Algebra/Order/Floor.lean +++ b/Mathlib/Algebra/Order/Floor.lean @@ -251,8 +251,6 @@ theorem preimage_floor_of_ne_zero {n : ℕ} (hn : n ≠ 0) : theorem lt_ceil : n < ⌈a⌉₊ ↔ (n : α) < a := lt_iff_lt_of_le_iff_le ceil_le --- porting note (#10618): simp can prove this --- @[simp] theorem add_one_le_ceil_iff : n + 1 ≤ ⌈a⌉₊ ↔ (n : α) < a := by rw [← Nat.lt_ceil, Nat.add_one_le_iff] @@ -924,8 +922,6 @@ theorem fract_ofNat (n : ℕ) [n.AtLeastTwo] : fract ((no_index (OfNat.ofNat n)) : α) = 0 := fract_natCast n --- porting note (#10618): simp can prove this --- @[simp] theorem fract_floor (a : α) : fract (⌊a⌋ : α) = 0 := fract_intCast _ diff --git a/Mathlib/Algebra/Order/Group/Unbundled/Basic.lean b/Mathlib/Algebra/Order/Group/Unbundled/Basic.lean index 95bd5c74cb92e..03c9d7579e5f1 100644 --- a/Mathlib/Algebra/Order/Group/Unbundled/Basic.lean +++ b/Mathlib/Algebra/Order/Group/Unbundled/Basic.lean @@ -148,7 +148,6 @@ theorem mul_inv_le_iff_le_mul : a * b⁻¹ ≤ c ↔ a ≤ c * b := theorem le_mul_inv_iff_mul_le : c ≤ a * b⁻¹ ↔ c * b ≤ a := (mul_le_mul_iff_right b).symm.trans <| by rw [inv_mul_cancel_right] --- Porting note (#10618): `simp` can prove this @[to_additive] theorem mul_inv_le_one_iff_le : a * b⁻¹ ≤ 1 ↔ a ≤ b := mul_inv_le_iff_le_mul.trans <| by rw [one_mul] @@ -193,7 +192,6 @@ theorem mul_inv_lt_iff_lt_mul : a * b⁻¹ < c ↔ a < c * b := by theorem lt_mul_inv_iff_mul_lt : c < a * b⁻¹ ↔ c * b < a := (mul_lt_mul_iff_right b).symm.trans <| by rw [inv_mul_cancel_right] --- Porting note (#10618): `simp` can prove this @[to_additive] theorem inv_mul_lt_one_iff_lt : a * b⁻¹ < 1 ↔ a < b := by rw [← mul_lt_mul_iff_right b, inv_mul_cancel_right, one_mul] diff --git a/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean b/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean index b0a55e98d4250..4001156b0e699 100644 --- a/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean +++ b/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean @@ -352,11 +352,9 @@ variable [LinearOrder α] {a b c : α} instance linearOrder : LinearOrder (WithZero α) := WithBot.linearOrder --- Porting note (#10618): @[simp] can prove this protected lemma le_max_iff : (a : WithZero α) ≤ max (b : WithZero α) c ↔ a ≤ max b c := by simp only [WithZero.coe_le_coe, le_max_iff] --- Porting note (#10618): @[simp] can prove this protected lemma min_le_iff : min (a : WithZero α) b ≤ c ↔ min a b ≤ c := by simp only [WithZero.coe_le_coe, min_le_iff] diff --git a/Mathlib/Algebra/Order/Monoid/Canonical/Defs.lean b/Mathlib/Algebra/Order/Monoid/Canonical/Defs.lean index f3c8c9e9daa85..b18b3c81d456e 100644 --- a/Mathlib/Algebra/Order/Monoid/Canonical/Defs.lean +++ b/Mathlib/Algebra/Order/Monoid/Canonical/Defs.lean @@ -226,12 +226,10 @@ theorem min_mul_distrib (a b c : α) : min a (b * c) = min a (min a b * min a c) theorem min_mul_distrib' (a b c : α) : min (a * b) c = min (min a c * min b c) c := by simpa [min_comm _ c] using min_mul_distrib c a b --- Porting note (#10618): no longer `@[simp]`, as `simp` can prove this. @[to_additive] theorem one_min (a : α) : min 1 a = 1 := min_eq_left (one_le a) --- Porting note (#10618): no longer `@[simp]`, as `simp` can prove this. @[to_additive] theorem min_one (a : α) : min a 1 = 1 := min_eq_right (one_le a) diff --git a/Mathlib/Algebra/Order/Monoid/Unbundled/WithTop.lean b/Mathlib/Algebra/Order/Monoid/Unbundled/WithTop.lean index f23cb20dc1b90..cd7326219233b 100644 --- a/Mathlib/Algebra/Order/Monoid/Unbundled/WithTop.lean +++ b/Mathlib/Algebra/Order/Monoid/Unbundled/WithTop.lean @@ -113,12 +113,8 @@ theorem add_eq_coe : | some a, ⊤, c => by simp | some a, some b, c => by norm_cast; simp --- Porting note (#10618): simp can already prove this. --- @[simp] theorem add_coe_eq_top_iff {x : WithTop α} {y : α} : x + y = ⊤ ↔ x = ⊤ := by simp --- Porting note (#10618): simp can already prove this. --- @[simp] theorem coe_add_eq_top_iff {y : WithTop α} : ↑x + y = ⊤ ↔ y = ⊤ := by simp theorem add_right_cancel_iff [IsRightCancelAdd α] (ha : a ≠ ⊤) : b + a = c + a ↔ b = c := by @@ -542,13 +538,9 @@ theorem bot_lt_add [LT α] {a b : WithBot α} : ⊥ < a + b ↔ ⊥ < a ∧ ⊥ theorem add_eq_coe : a + b = x ↔ ∃ a' b' : α, ↑a' = a ∧ ↑b' = b ∧ a' + b' = x := WithTop.add_eq_coe --- Porting note (#10618): simp can already prove this. --- @[simp] theorem add_coe_eq_bot_iff : a + y = ⊥ ↔ a = ⊥ := WithTop.add_coe_eq_top_iff --- Porting note (#10618): simp can already prove this. --- @[simp] theorem coe_add_eq_bot_iff : ↑x + b = ⊥ ↔ b = ⊥ := WithTop.coe_add_eq_top_iff diff --git a/Mathlib/Algebra/Polynomial/Basic.lean b/Mathlib/Algebra/Polynomial/Basic.lean index f095ae814ba7d..57f4bba5519be 100644 --- a/Mathlib/Algebra/Polynomial/Basic.lean +++ b/Mathlib/Algebra/Polynomial/Basic.lean @@ -455,7 +455,6 @@ theorem smul_C {S} [SMulZeroClass S R] (s : S) (r : R) : s • C r = C (s • r) theorem C_pow : C (a ^ n) = C a ^ n := C.map_pow a n --- @[simp] -- Porting note (#10618): simp can prove this theorem C_eq_natCast (n : ℕ) : C (n : R) = (n : R[X]) := map_natCast C n diff --git a/Mathlib/Algebra/Polynomial/Coeff.lean b/Mathlib/Algebra/Polynomial/Coeff.lean index 82ae37b47a959..383ce0d93d741 100644 --- a/Mathlib/Algebra/Polynomial/Coeff.lean +++ b/Mathlib/Algebra/Polynomial/Coeff.lean @@ -346,7 +346,7 @@ theorem natCast_coeff_zero {n : ℕ} {R : Type*} [Semiring R] : (n : R[X]).coeff @[deprecated (since := "2024-04-17")] alias nat_cast_coeff_zero := natCast_coeff_zero -@[norm_cast] -- @[simp] -- Porting note (#10618): simp can prove this +@[norm_cast] theorem natCast_inj {m n : ℕ} {R : Type*} [Semiring R] [CharZero R] : (↑m : R[X]) = ↑n ↔ m = n := by constructor @@ -366,7 +366,7 @@ theorem intCast_coeff_zero {i : ℤ} {R : Type*} [Ring R] : (i : R[X]).coeff 0 = @[deprecated (since := "2024-04-17")] alias int_cast_coeff_zero := intCast_coeff_zero -@[norm_cast] -- @[simp] -- Porting note (#10618): simp can prove this +@[norm_cast] theorem intCast_inj {m n : ℤ} {R : Type*} [Ring R] [CharZero R] : (↑m : R[X]) = ↑n ↔ m = n := by constructor · intro h diff --git a/Mathlib/Algebra/Polynomial/Degree/Definitions.lean b/Mathlib/Algebra/Polynomial/Degree/Definitions.lean index 9ec69895a9634..e5cc71815055f 100644 --- a/Mathlib/Algebra/Polynomial/Degree/Definitions.lean +++ b/Mathlib/Algebra/Polynomial/Degree/Definitions.lean @@ -758,11 +758,9 @@ theorem leadingCoeff_C_mul_X (a : R) : leadingCoeff (C a * X) = a := by theorem leadingCoeff_C (a : R) : leadingCoeff (C a) = a := leadingCoeff_monomial a 0 --- @[simp] -- Porting note (#10618): simp can prove this theorem leadingCoeff_X_pow (n : ℕ) : leadingCoeff ((X : R[X]) ^ n) = 1 := by simpa only [C_1, one_mul] using leadingCoeff_C_mul_X_pow (1 : R) n --- @[simp] -- Porting note (#10618): simp can prove this theorem leadingCoeff_X : leadingCoeff (X : R[X]) = 1 := by simpa only [pow_one] using @leadingCoeff_X_pow R _ 1 @@ -774,7 +772,6 @@ theorem monic_X_pow (n : ℕ) : Monic (X ^ n : R[X]) := theorem monic_X : Monic (X : R[X]) := leadingCoeff_X --- @[simp] -- Porting note (#10618): simp can prove this theorem leadingCoeff_one : leadingCoeff (1 : R[X]) = 1 := leadingCoeff_C 1 diff --git a/Mathlib/Algebra/Polynomial/Derivative.lean b/Mathlib/Algebra/Polynomial/Derivative.lean index 5e8f7007d1c5b..899b80f020a97 100644 --- a/Mathlib/Algebra/Polynomial/Derivative.lean +++ b/Mathlib/Algebra/Polynomial/Derivative.lean @@ -98,7 +98,6 @@ theorem derivative_C_mul_X_sq (a : R) : derivative (C a * X ^ 2) = C (a * 2) * X theorem derivative_X_pow (n : ℕ) : derivative (X ^ n : R[X]) = C (n : R) * X ^ (n - 1) := by convert derivative_C_mul_X_pow (1 : R) n <;> simp --- Porting note (#10618): removed `simp`: `simp` can prove it. theorem derivative_X_sq : derivative (X ^ 2 : R[X]) = C 2 * X := by rw [derivative_X_pow, Nat.cast_two, pow_one] @@ -120,16 +119,13 @@ theorem derivative_one : derivative (1 : R[X]) = 0 := theorem derivative_add {f g : R[X]} : derivative (f + g) = derivative f + derivative g := derivative.map_add f g --- Porting note (#10618): removed `simp`: `simp` can prove it. theorem derivative_X_add_C (c : R) : derivative (X + C c) = 1 := by rw [derivative_add, derivative_X, derivative_C, add_zero] --- Porting note (#10618): removed `simp`: `simp` can prove it. theorem derivative_sum {s : Finset ι} {f : ι → R[X]} : derivative (∑ b ∈ s, f b) = ∑ b ∈ s, derivative (f b) := map_sum .. --- Porting note (#10618): removed `simp`: `simp` can prove it. theorem derivative_smul {S : Type*} [Monoid S] [DistribMulAction S R] [IsScalarTower S R R] (s : S) (p : R[X]) : derivative (s • p) = s • derivative p := derivative.map_smul_of_tower s p @@ -571,7 +567,6 @@ theorem iterate_derivative_neg {f : R[X]} {k : ℕ} : derivative^[k] (-f) = -der theorem derivative_sub {f g : R[X]} : derivative (f - g) = derivative f - derivative g := LinearMap.map_sub derivative f g --- Porting note (#10618): removed `simp`: `simp` can prove it. theorem derivative_X_sub_C (c : R) : derivative (X - C c) = 1 := by rw [derivative_sub, derivative_X, derivative_C, sub_zero] diff --git a/Mathlib/Algebra/Polynomial/Laurent.lean b/Mathlib/Algebra/Polynomial/Laurent.lean index 7210a0ec6104f..52e9ac7978fee 100644 --- a/Mathlib/Algebra/Polynomial/Laurent.lean +++ b/Mathlib/Algebra/Polynomial/Laurent.lean @@ -218,7 +218,6 @@ theorem _root_.Polynomial.toLaurent_C_mul_eq (r : R) (f : R[X]) : theorem _root_.Polynomial.toLaurent_X_pow (n : ℕ) : toLaurent (X ^ n : R[X]) = T n := by simp only [map_pow, Polynomial.toLaurent_X, T_pow, mul_one] --- @[simp] -- Porting note (#10618): simp can prove this theorem _root_.Polynomial.toLaurent_C_mul_X_pow (n : ℕ) (r : R) : toLaurent (Polynomial.C r * X ^ n) = C r * T n := by simp only [_root_.map_mul, Polynomial.toLaurent_C, Polynomial.toLaurent_X_pow] diff --git a/Mathlib/Algebra/Polynomial/Reverse.lean b/Mathlib/Algebra/Polynomial/Reverse.lean index 614061ef26c6d..2d9684409878b 100644 --- a/Mathlib/Algebra/Polynomial/Reverse.lean +++ b/Mathlib/Algebra/Polynomial/Reverse.lean @@ -78,7 +78,6 @@ theorem revAt_add {N O n o : ℕ} (hn : n ≤ N) (ho : o ≤ O) : rw [add_assoc, add_left_comm n' o, ← add_assoc, revAt_le (le_add_right rfl.le)] repeat' rw [add_tsub_cancel_left] --- @[simp] -- Porting note (#10618): simp can prove this theorem revAt_zero (N : ℕ) : revAt N 0 = N := by simp /-- `reflect N f` is the polynomial such that `(reflect N f).coeff i = f.coeff (revAt N i)`. @@ -123,7 +122,6 @@ theorem reflect_C_mul (f : R[X]) (r : R) (N : ℕ) : reflect N (C r * f) = C r * ext simp only [coeff_reflect, coeff_C_mul] --- @[simp] -- Porting note (#10618): simp can prove this (once `reflect_monomial` is in simp scope) theorem reflect_C_mul_X_pow (N n : ℕ) {c : R} : reflect N (C c * X ^ n) = C c * X ^ revAt N n := by ext rw [reflect_C_mul, coeff_C_mul, coeff_C_mul, coeff_X_pow, coeff_reflect] diff --git a/Mathlib/Algebra/Quaternion.lean b/Mathlib/Algebra/Quaternion.lean index 05c61e0d4e530..6fe9d273656aa 100644 --- a/Mathlib/Algebra/Quaternion.lean +++ b/Mathlib/Algebra/Quaternion.lean @@ -1282,15 +1282,12 @@ instance instDivisionRing : DivisionRing ℍ[R] where nnqsmul_def q x := by rw [← coe_nnratCast, coe_mul_eq_smul]; ext <;> exact NNRat.smul_def _ _ qsmul_def q x := by rw [← coe_ratCast, coe_mul_eq_smul]; ext <;> exact Rat.smul_def _ _ ---@[simp] Porting note (#10618): `simp` can prove it theorem normSq_inv : normSq a⁻¹ = (normSq a)⁻¹ := map_inv₀ normSq _ ---@[simp] Porting note (#10618): `simp` can prove it theorem normSq_div : normSq (a / b) = normSq a / normSq b := map_div₀ normSq a b ---@[simp] Porting note (#10618): `simp` can prove it theorem normSq_zpow (z : ℤ) : normSq (a ^ z) = normSq a ^ z := map_zpow₀ normSq a z @@ -1330,7 +1327,6 @@ theorem mk_quaternionAlgebra_of_infinite [Infinite R] : #(ℍ[R,c₁,c₂]) = #R theorem mk_univ_quaternionAlgebra : #(Set.univ : Set ℍ[R,c₁,c₂]) = #R ^ 4 := by rw [mk_univ, mk_quaternionAlgebra] ---@[simp] Porting note (#10618): `simp` can prove it theorem mk_univ_quaternionAlgebra_of_infinite [Infinite R] : #(Set.univ : Set ℍ[R,c₁,c₂]) = #R := by rw [mk_univ_quaternionAlgebra, pow_four] @@ -1359,7 +1355,6 @@ theorem mk_quaternion_of_infinite [Infinite R] : #(ℍ[R]) = #R := mk_quaternionAlgebra_of_infinite _ _ /-- The cardinality of the quaternions, as a set. -/ ---@[simp] Porting note (#10618): `simp` can prove it theorem mk_univ_quaternion : #(Set.univ : Set ℍ[R]) = #R ^ 4 := mk_univ_quaternionAlgebra _ _ diff --git a/Mathlib/Algebra/Ring/BooleanRing.lean b/Mathlib/Algebra/Ring/BooleanRing.lean index 985d2359dfa85..e88f282f73822 100644 --- a/Mathlib/Algebra/Ring/BooleanRing.lean +++ b/Mathlib/Algebra/Ring/BooleanRing.lean @@ -137,11 +137,9 @@ theorem toBoolAlg_ofBoolAlg (a : AsBoolAlg α) : toBoolAlg (ofBoolAlg a) = a := theorem ofBoolAlg_toBoolAlg (a : α) : ofBoolAlg (toBoolAlg a) = a := rfl --- Porting note (#10618): simp can prove this -- @[simp] theorem toBoolAlg_inj {a b : α} : toBoolAlg a = toBoolAlg b ↔ a = b := Iff.rfl --- Porting note (#10618): simp can prove this -- @[simp] theorem ofBoolAlg_inj {a b : AsBoolAlg α} : ofBoolAlg a = ofBoolAlg b ↔ a = b := Iff.rfl @@ -355,11 +353,9 @@ theorem toBoolRing_ofBoolRing (a : AsBoolRing α) : toBoolRing (ofBoolRing a) = theorem ofBoolRing_toBoolRing (a : α) : ofBoolRing (toBoolRing a) = a := rfl --- Porting note (#10618): simp can prove this -- @[simp] theorem toBoolRing_inj {a b : α} : toBoolRing a = toBoolRing b ↔ a = b := Iff.rfl --- Porting note (#10618): simp can prove this -- @[simp] theorem ofBoolRing_inj {a b : AsBoolRing α} : ofBoolRing a = ofBoolRing b ↔ a = b := Iff.rfl diff --git a/Mathlib/Algebra/Ring/Commute.lean b/Mathlib/Algebra/Ring/Commute.lean index 27887a07dc5bb..381b7109ee286 100644 --- a/Mathlib/Algebra/Ring/Commute.lean +++ b/Mathlib/Algebra/Ring/Commute.lean @@ -77,13 +77,9 @@ section variable [MulOneClass R] [HasDistribNeg R] --- Porting note (#10618): no longer needs to be `@[simp]` since `simp` can prove it. --- @[simp] theorem neg_one_right (a : R) : Commute a (-1) := SemiconjBy.neg_one_right a --- Porting note (#10618): no longer needs to be `@[simp]` since `simp` can prove it. --- @[simp] theorem neg_one_left (a : R) : Commute (-1) a := SemiconjBy.neg_one_left a diff --git a/Mathlib/Algebra/Ring/CompTypeclasses.lean b/Mathlib/Algebra/Ring/CompTypeclasses.lean index 7ab747f493635..6aff4dd2a3908 100644 --- a/Mathlib/Algebra/Ring/CompTypeclasses.lean +++ b/Mathlib/Algebra/Ring/CompTypeclasses.lean @@ -81,22 +81,16 @@ class RingHomInvPair (σ : R₁ →+* R₂) (σ' : outParam (R₂ →+* R₁)) : /-- `σ'` is a left inverse of `σ'` -/ comp_eq₂ : σ.comp σ' = RingHom.id R₂ --- attribute [simp] RingHomInvPair.comp_eq Porting note (#10618): `simp` can prove it - --- attribute [simp] RingHomInvPair.comp_eq₂ Porting note (#10618): `simp` can prove it - variable {σ : R₁ →+* R₂} {σ' : R₂ →+* R₁} namespace RingHomInvPair variable [RingHomInvPair σ σ'] --- @[simp] Porting note (#10618): `simp` can prove it theorem comp_apply_eq {x : R₁} : σ' (σ x) = x := by rw [← RingHom.comp_apply, comp_eq] simp --- @[simp] Porting note (#10618): `simp` can prove it theorem comp_apply_eq₂ {x : R₂} : σ (σ' x) = x := by rw [← RingHom.comp_apply, comp_eq₂] simp diff --git a/Mathlib/Algebra/Ring/Equiv.lean b/Mathlib/Algebra/Ring/Equiv.lean index 8518ce1badcb4..ff0a1107f04c7 100644 --- a/Mathlib/Algebra/Ring/Equiv.lean +++ b/Mathlib/Algebra/Ring/Equiv.lean @@ -733,12 +733,10 @@ theorem toMonoidHom_refl : (RingEquiv.refl R).toMonoidHom = MonoidHom.id R := theorem toAddMonoidHom_refl : (RingEquiv.refl R).toAddMonoidHom = AddMonoidHom.id R := rfl --- Porting note (#10618): Now other `simp` can do this, so removed `simp` attribute theorem toRingHom_apply_symm_toRingHom_apply (e : R ≃+* S) : ∀ y : S, e.toRingHom (e.symm.toRingHom y) = y := e.toEquiv.apply_symm_apply --- Porting note (#10618): Now other `simp` can do this, so removed `simp` attribute theorem symm_toRingHom_apply_toRingHom_apply (e : R ≃+* S) : ∀ x : R, e.symm.toRingHom (e.toRingHom x) = x := Equiv.symm_apply_apply e.toEquiv @@ -748,13 +746,11 @@ theorem toRingHom_trans (e₁ : R ≃+* S) (e₂ : S ≃+* S') : (e₁.trans e₂).toRingHom = e₂.toRingHom.comp e₁.toRingHom := rfl --- Porting note (#10618): Now other `simp` can do this, so removed `simp` attribute theorem toRingHom_comp_symm_toRingHom (e : R ≃+* S) : e.toRingHom.comp e.symm.toRingHom = RingHom.id _ := by ext simp --- Porting note (#10618): Now other `simp` can do this, so removed `simp` attribute theorem symm_toRingHom_comp_toRingHom (e : R ≃+* S) : e.symm.toRingHom.comp e.toRingHom = RingHom.id _ := by ext diff --git a/Mathlib/Algebra/Ring/Parity.lean b/Mathlib/Algebra/Ring/Parity.lean index b4890ab73e4bc..72fdec393def4 100644 --- a/Mathlib/Algebra/Ring/Parity.lean +++ b/Mathlib/Algebra/Ring/Parity.lean @@ -168,11 +168,6 @@ end Monoid section Ring variable [Ring α] {a b : α} {n : ℕ} -/- Porting note (#10618): attribute `simp` removed based on linter report -simp can prove this: - by simp only [even_neg, even_two] --/ --- @[simp] lemma even_neg_two : Even (-2 : α) := by simp only [even_neg, even_two] lemma Odd.neg (hp : Odd a) : Odd (-a) := by @@ -183,11 +178,6 @@ lemma Odd.neg (hp : Odd a) : Odd (-a) := by @[simp] lemma odd_neg : Odd (-a) ↔ Odd a := ⟨fun h ↦ neg_neg a ▸ h.neg, Odd.neg⟩ -/- Porting note (#10618): attribute `simp` removed based on linter report -simp can prove this: - by simp only [odd_neg, odd_one] --/ --- @[simp] lemma odd_neg_one : Odd (-1 : α) := by simp lemma Odd.sub_even (ha : Odd a) (hb : Even b) : Odd (a - b) := by diff --git a/Mathlib/Algebra/Ring/Subring/Basic.lean b/Mathlib/Algebra/Ring/Subring/Basic.lean index e16be89dd2957..35a508fb82fb3 100644 --- a/Mathlib/Algebra/Ring/Subring/Basic.lean +++ b/Mathlib/Algebra/Ring/Subring/Basic.lean @@ -358,8 +358,6 @@ theorem coe_one : ((1 : s) : R) = 1 := theorem coe_pow (x : s) (n : ℕ) : ↑(x ^ n) = (x : R) ^ n := SubmonoidClass.coe_pow x n --- TODO: can be generalized to `AddSubmonoidClass` --- @[simp] -- Porting note (#10618): simp can prove this theorem coe_eq_zero_iff {x : s} : (x : R) = 0 ↔ x = 0 := ⟨fun h => Subtype.ext (Trans.trans h s.coe_zero.symm), fun h => h.symm ▸ s.coe_zero⟩ @@ -1100,11 +1098,9 @@ def inclusion {S T : Subring R} (h : S ≤ T) : S →+* T := theorem range_subtype (s : Subring R) : s.subtype.range = s := SetLike.coe_injective <| (coe_rangeS _).trans Subtype.range_coe --- @[simp] -- Porting note (#10618): simp can prove this theorem range_fst : (fst R S).rangeS = ⊤ := (fst R S).rangeS_top_of_surjective <| Prod.fst_surjective --- @[simp] -- Porting note (#10618): simp can prove this theorem range_snd : (snd R S).rangeS = ⊤ := (snd R S).rangeS_top_of_surjective <| Prod.snd_surjective diff --git a/Mathlib/Algebra/Ring/Subsemiring/Basic.lean b/Mathlib/Algebra/Ring/Subsemiring/Basic.lean index a67f00afb054d..589f27314eac8 100644 --- a/Mathlib/Algebra/Ring/Subsemiring/Basic.lean +++ b/Mathlib/Algebra/Ring/Subsemiring/Basic.lean @@ -152,7 +152,6 @@ instance : SubsemiringClass (Subsemiring R) R where theorem mem_toSubmonoid {s : Subsemiring R} {x : R} : x ∈ s.toSubmonoid ↔ x ∈ s := Iff.rfl --- `@[simp]` -- Porting note (#10618): simp can prove thisrove this theorem mem_carrier {s : Subsemiring R} {x : R} : x ∈ s.carrier ↔ x ∈ s := Iff.rfl diff --git a/Mathlib/Algebra/Squarefree/Basic.lean b/Mathlib/Algebra/Squarefree/Basic.lean index 935cf00ec3b26..c18ad3b0214ad 100644 --- a/Mathlib/Algebra/Squarefree/Basic.lean +++ b/Mathlib/Algebra/Squarefree/Basic.lean @@ -42,7 +42,6 @@ theorem IsRelPrime.of_squarefree_mul [CommMonoid R] {m n : R} (h : Squarefree (m theorem IsUnit.squarefree [CommMonoid R] {x : R} (h : IsUnit x) : Squarefree x := fun _ hdvd => isUnit_of_mul_isUnit_left (isUnit_of_dvd_unit hdvd h) --- @[simp] -- Porting note (#10618): simp can prove this theorem squarefree_one [CommMonoid R] : Squarefree (1 : R) := isUnit_one.squarefree diff --git a/Mathlib/Algebra/Symmetrized.lean b/Mathlib/Algebra/Symmetrized.lean index f4dea0e4e8345..66841a63dca0e 100644 --- a/Mathlib/Algebra/Symmetrized.lean +++ b/Mathlib/Algebra/Symmetrized.lean @@ -95,11 +95,9 @@ theorem unsym_injective : Injective (unsym : αˢʸᵐ → α) := theorem unsym_surjective : Surjective (unsym : αˢʸᵐ → α) := unsym.surjective --- Porting note (#10618): @[simp] can prove this theorem sym_inj {a b : α} : sym a = sym b ↔ a = b := sym_injective.eq_iff --- Porting note (#10618): @[simp] can prove this theorem unsym_inj {a b : αˢʸᵐ} : unsym a = unsym b ↔ a = b := unsym_injective.eq_iff diff --git a/Mathlib/Algebra/Tropical/Basic.lean b/Mathlib/Algebra/Tropical/Basic.lean index 373542c05d885..5a22f3bbfe005 100644 --- a/Mathlib/Algebra/Tropical/Basic.lean +++ b/Mathlib/Algebra/Tropical/Basic.lean @@ -291,7 +291,6 @@ theorem add_eq_left_iff {x y : Tropical R} : x + y = x ↔ x ≤ y := by theorem add_eq_right_iff {x y : Tropical R} : x + y = y ↔ y ≤ x := by rw [trop_add_def, trop_eq_iff_eq_untrop, ← untrop_le_iff, min_eq_right_iff] --- Porting note (#10618): removing `simp`. `simp` can prove it theorem add_self (x : Tropical R) : x + x = x := untrop_injective (min_eq_right le_rfl) @@ -495,7 +494,6 @@ theorem succ_nsmul {R} [LinearOrder R] [OrderTop R] (x : Tropical R) (n : ℕ) : -- Requires `zero_eq_bot` to be true -- lemma add_eq_zero_iff {a b : tropical R} : -- a + b = 1 ↔ a = 1 ∨ b = 1 := sorry --- Porting note (#10618): removing @[simp], `simp` can prove it theorem mul_eq_zero_iff {R : Type*} [LinearOrderedAddCommMonoid R] {a b : Tropical (WithTop R)} : a * b = 0 ↔ a = 0 ∨ b = 0 := by simp [← untrop_inj_iff, WithTop.add_eq_top] From fab5497a2d3ee1fd2526d57093abedf3ffc12f54 Mon Sep 17 00:00:00 2001 From: damiano Date: Fri, 25 Oct 2024 09:18:09 +0000 Subject: [PATCH 408/505] feat: `Deprecated` is the new `Init` for tech-debts (#18090) There are no more `Init` files, so this PR * removes `Init` counts from the `tech-debt` counters; * uses the same fine-grained approach that `Init` had for `Deprecated`. I also wrote a better script to compute differences that should address the issues that the previous code had. I simply removed the old test, rather than muting it, since I do not imagine anyone adding new `Init` files. --- scripts/technical-debt-metrics.sh | 78 +++++++++++++++++-------------- 1 file changed, 44 insertions(+), 34 deletions(-) diff --git a/scripts/technical-debt-metrics.sh b/scripts/technical-debt-metrics.sh index ce9a3a5202b16..524c97f135331 100755 --- a/scripts/technical-debt-metrics.sh +++ b/scripts/technical-debt-metrics.sh @@ -15,6 +15,28 @@ IFS=$'\n\t' currCommit="${1:-"$(git rev-parse HEAD)"}" refCommit="${2:-"$(git log --pretty=%H --since="$(date -I -d 'last week')" | tail -n -1)"}" +## `computeDiff input` assumes that input consists of lines of the form `value|description` +## where `value` is a number and `description` is the statistic that `value` reports. +## `value` is non-negative, if it refers to the current commit and it is negative otherwise. +## The output is of the form `|current value|difference|description|`. +computeDiff () { + awk -F'|' 'BEGIN{con=1}{ + # order keeps track of maintaining the final order of the counters the same as the input one + # rdict stores the difference of the current value minus the old value + # curr stores the current value, making sure that the number is non-negative and does not start with `-` + if(rdict[$2] == "") {order[con]=$2; con++} + if((0 <= $1+0) && (!($1 ~ "-"))) {curr[$2]=$1} + rdict[$2]+=$1+0 + } END { + # loop over the "sorted" index in `order` + for(ind=1; ind<=con-1; ind++) { + # retrieve the description of the counter + val=order[ind] + # print `|current value|difference|name of counter|` + printf("|%s|%s|%s|\n", curr[val], rdict[val], val)} + }' "${1}" +} + # `tdc` produces a semi-formatted output of the form # ... # |description @@ -34,7 +56,6 @@ titlesAndRegexes=( "maxHeartBeats modifications" "^ *set_option .*maxHeartbeats" ) -printf '|Current number|Change|Type|\n|-:|:-:|:-|\n' for i in ${!titlesAndRegexes[@]}; do # loop on the odd-indexed entries and name each entry and the following if (( (i + 1) % 2 )); then @@ -58,43 +79,32 @@ printf '%s|%s\n' "$(wc -l < scripts/no_lints_prime_decls.txt)" "exceptions for t deprecatedFiles="$(git ls-files '**/Deprecated/*.lean' | xargs wc -l | sed 's=^ *==')" printf '%s|%s\n' "$(printf '%s' "${deprecatedFiles}" | wc -l)" "\`Deprecated\` files" -printf '%s|%s\n' "$(printf '%s\n' "${deprecatedFiles}" | grep total | sed 's= total==')" 'total LoC in `Deprecated` files' - -initFiles="$(git ls-files '**/Init/*.lean' | xargs wc -l | sed 's=^ *==')" - -printf '%s|%s\n' "$(printf '%s' "${initFiles}" | wc -l)" "\`Init\` files" -printf '%s|%s\n\n' "$(printf '%s\n' "${initFiles}" | grep total | sed 's= total==')" 'total LoC in `Init` files' - -printf $'```spoiler Changed \'Init\' lines by file\n%s\n```\n' "$( - printf '%s\n' "${initFiles}" | awk 'BEGIN{print("|LoC|Change|File|\n|-:|:-:|-|")} {printf("%s|%s\n", $1, $2)}' - )" +printf '%s|%s\n\n' "$(printf '%s\n' "${deprecatedFiles}" | grep total | sed 's= total==')" 'total LoC in `Deprecated` files' } -# collect the technical debts from the current mathlib -new="$(git checkout -q "${currCommit}" && tdc; git switch -q -)" +# collect the technical debts and the line counts of the deprecated file from the current mathlib +git checkout -q "${currCommit}" +new="$(tdc)" +newDeprecatedFiles="$(git ls-files '**/Deprecated/*.lean' | xargs wc -l | sed 's=^ *==')" +git switch -q - -# collect the technical debts from the reference mathlib -- switch to the -old="$(git checkout -q "${refCommit}" && tdc; git switch -q -)" +# collect the technical debts and the line counts of the deprecated file from the reference mathlib +git checkout -q "${refCommit}" +old="$(tdc | sed 's=^[0-9]=-&=')" +oldDeprecatedFiles="$(git ls-files '**/Deprecated/*.lean' | xargs wc -l | sed 's=^ *=-=')" +git switch -q - -# place the outputs side-by-side, using `@` as a separator -paste -d@ <(echo "$new") <(echo "${old}") | sed 's=@=|@|=' | - # each line should look like eithe - # [new number]|description@[old number]|descr - # or something that does not start with [number]| - # we split the lines into "words", using `|` as a separator - awk -F'|' ' - # if the first "word" is a number, then we write the 1st entry and compare it with the 4th - ($1+0 == $1) { printf("|%s|%s|%s|\n", $1, $1-$4, $2) } - # otherwise, the line is a "formatting" line, so we simply print it unchanged until we find `@` - !($1+0 == $1) { - for(i=1; i<=NF; i++) { - if ($i == "@") { break } - else { printf("%s|", $i) } - } - print "@" - }' | - # the sequence `@|` ending a line is an artifact of our process and we remove it - sed 's=|@$==' +printf '|Current number|Change|Type|\n|-:|:-:|:-|\n' +printf '%s\n%s\n' "${old}" "${new}" | computeDiff - +deprSummary="$(printf '%s\n%s\n' "${oldDeprecatedFiles}" "${newDeprecatedFiles}" | tr ' ' '|' | computeDiff -)" + +printf $'```spoiler Changed \'Deprecated\' lines by file\n%s\n```\n' "$( + printf '%s\n' "${deprSummary}" | awk -F'|' 'BEGIN{print("|LoC|Change|File|\n|-:|:-:|-|")} + ($4 == "total") {total=$0} + (!($4 == "total")) { + printf("%s\n", $0) + } END {printf("%s\n", total)}' + )" baseURL='https://github.com/leanprover-community/mathlib4/commit' printf '\nCurrent commit [%s](%s)\n' "${currCommit:0:10}" "${baseURL}/${currCommit}" From 8e419e3d18cb1b1aaf91069a76545bd3269ee69b Mon Sep 17 00:00:00 2001 From: Scott Carnahan <128885296+ScottCarnahan@users.noreply.github.com> Date: Fri, 25 Oct 2024 09:55:06 +0000 Subject: [PATCH 409/505] feat (Data/Finset/Basic) : Equivalence between Pi type over Finset.cons and a product (#17004) This PR defines an equivalence between Pi over an augmented finset and a product with Pi over the original finset. --- Mathlib/Data/Finset/Basic.lean | 70 ++++++++++++++++++- .../GroupTheory/SpecificGroups/KleinFour.lean | 4 +- Mathlib/Order/CountableDenseLinearOrder.lean | 6 +- 3 files changed, 73 insertions(+), 7 deletions(-) diff --git a/Mathlib/Data/Finset/Basic.lean b/Mathlib/Data/Finset/Basic.lean index 7f0708bdd080c..e3a1af524b590 100644 --- a/Mathlib/Data/Finset/Basic.lean +++ b/Mathlib/Data/Finset/Basic.lean @@ -751,6 +751,13 @@ theorem mem_cons_self (a : α) (s : Finset α) {h} : a ∈ cons a s h := theorem cons_val (h : a ∉ s) : (cons a s h).1 = a ::ₘ s.1 := rfl +theorem eq_of_mem_cons_of_not_mem (has : a ∉ s) (h : b ∈ cons a s has) (hb : b ∉ s) : b = a := + (mem_cons.1 h).resolve_right hb + +theorem mem_of_mem_cons_of_ne {s : Finset α} {a : α} {has} {i : α} + (hi : i ∈ cons a s has) (hia : i ≠ a) : i ∈ s := + (mem_cons.1 hi).resolve_left hia + theorem forall_mem_cons (h : a ∉ s) (p : α → Prop) : (∀ x, x ∈ cons a s h → p x) ↔ p a ∧ ∀ x, x ∈ s → p x := by simp only [mem_cons, or_imp, forall_and, forall_eq] @@ -808,6 +815,36 @@ theorem cons_swap (hb : b ∉ s) (ha : a ∉ s.cons b hb) : 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 +/-- Split the added element of cons off a Pi type. -/ +@[simps!] +def consPiProd (f : α → Type*) (has : a ∉ s) (x : Π i ∈ cons a s has, f i) : f a × Π i ∈ s, f i := + (x a (mem_cons_self a s), fun i hi => x i (mem_cons_of_mem hi)) + +/-- Combine a product with a pi type to pi of cons. -/ +def prodPiCons [DecidableEq α] (f : α → Type*) {a : α} (has : a ∉ s) (x : f a × Π i ∈ s, f i) : + (Π i ∈ cons a s has, f i) := + fun i hi => + if h : i = a then cast (congrArg f h.symm) x.1 else x.2 i (mem_of_mem_cons_of_ne hi h) + +/-- The equivalence between pi types on cons and the product. -/ +def consPiProdEquiv [DecidableEq α] {s : Finset α} (f : α → Type*) {a : α} (has : a ∉ s) : + (Π i ∈ cons a s has, f i) ≃ f a × Π i ∈ s, f i where + toFun := consPiProd f has + invFun := prodPiCons f has + left_inv _ := by + ext i _ + dsimp only [prodPiCons, consPiProd] + by_cases h : i = a + · rw [dif_pos h] + subst h + simp_all only [cast_eq] + · rw [dif_neg h] + right_inv _ := by + ext _ hi + · simp [prodPiCons] + · simp only [consPiProd_snd] + exact dif_neg (ne_of_mem_of_not_mem hi has) + end Cons /-! ### disjoint -/ @@ -955,7 +992,7 @@ theorem mem_insert_of_mem (h : a ∈ s) : a ∈ insert b s := theorem mem_of_mem_insert_of_ne (h : b ∈ insert a s) : b ≠ a → b ∈ s := (mem_insert.1 h).resolve_left -theorem eq_of_not_mem_of_mem_insert (ha : b ∈ insert a s) (hb : b ∉ s) : b = a := +theorem eq_of_mem_insert_of_not_mem (ha : b ∈ insert a s) (hb : b ∉ s) : b = a := (mem_insert.1 ha).resolve_right hb /-- A version of `LawfulSingleton.insert_emptyc_eq` that works with `dsimp`. -/ @@ -1039,7 +1076,7 @@ theorem insert_subset_insert (a : α) {s t : Finset α} (h : s ⊆ t) : insert a simp_rw [← coe_subset]; simp [-coe_subset, ha] theorem insert_inj (ha : a ∉ s) : insert a s = insert b s ↔ a = b := - ⟨fun h => eq_of_not_mem_of_mem_insert (h ▸ mem_insert_self _ _) ha, congr_arg (insert · s)⟩ + ⟨fun h => eq_of_mem_insert_of_not_mem (h ▸ mem_insert_self _ _) ha, congr_arg (insert · s)⟩ theorem insert_inj_on (s : Finset α) : Set.InjOn (fun a => insert a s) sᶜ := fun _ h _ _ => (insert_inj h).1 @@ -1137,6 +1174,35 @@ theorem disjoint_insert_left : Disjoint (insert a s) t ↔ a ∉ t ∧ Disjoint theorem disjoint_insert_right : Disjoint s (insert a t) ↔ a ∉ s ∧ Disjoint s t := disjoint_comm.trans <| by rw [disjoint_insert_left, _root_.disjoint_comm] +/-- Split the added element of insert off a Pi type. -/ +@[simps!] +def insertPiProd (f : α → Type*) (x : Π i ∈ insert a s, f i) : f a × Π i ∈ s, f i := + (x a (mem_insert_self a s), fun i hi => x i (mem_insert_of_mem hi)) + +/-- Combine a product with a pi type to pi of insert. -/ +def prodPiInsert (f : α → Type*) {a : α} (x : f a × Π i ∈ s, f i) : (Π i ∈ insert a s, f i) := + fun i hi => + if h : i = a then cast (congrArg f h.symm) x.1 else x.2 i (mem_of_mem_insert_of_ne hi h) + +/-- The equivalence between pi types on insert and the product. -/ +def insertPiProdEquiv [DecidableEq α] {s : Finset α} (f : α → Type*) {a : α} (has : a ∉ s) : + (Π i ∈ insert a s, f i) ≃ f a × Π i ∈ s, f i where + toFun := insertPiProd f + invFun := prodPiInsert f + left_inv _ := by + ext i _ + dsimp only [prodPiInsert, insertPiProd] + by_cases h : i = a + · rw [dif_pos h] + subst h + simp_all only [cast_eq] + · rw [dif_neg h] + right_inv _ := by + ext _ hi + · simp [prodPiInsert] + · simp only [insertPiProd_snd] + exact dif_neg (ne_of_mem_of_not_mem hi has) + end Insert /-! ### Lattice structure -/ diff --git a/Mathlib/GroupTheory/SpecificGroups/KleinFour.lean b/Mathlib/GroupTheory/SpecificGroups/KleinFour.lean index b902f89c46909..e5a725fd32218 100644 --- a/Mathlib/GroupTheory/SpecificGroups/KleinFour.lean +++ b/Mathlib/GroupTheory/SpecificGroups/KleinFour.lean @@ -120,7 +120,7 @@ lemma eq_mul_of_ne_all {x y z : G} (hx : x ≠ 1) (hy : y ≠ 1) (hxy : x ≠ y) (hz : z ≠ 1) (hzx : z ≠ x) (hzy : z ≠ y) : z = x * y := by classical let _ := Fintype.ofFinite G - apply eq_of_not_mem_of_mem_insert <| (eq_finset_univ hx hy hxy).symm ▸ mem_univ _ + apply eq_of_mem_insert_of_not_mem <| (eq_finset_univ hx hy hxy).symm ▸ mem_univ _ simpa only [mem_singleton, mem_insert, not_or] using ⟨hzx, hzy, hz⟩ variable {G₁ G₂ : Type*} [Group G₁] [Group G₂] [IsKleinFour G₁] @@ -146,7 +146,7 @@ def mulEquiv' (e : G₁ ≃ G₂) (he : e 1 = 1) (h : Monoid.exponent G₂ = 2) rw [← Ne, ← e.injective.ne_iff] at hx hy hxy rw [he] at hx hy symm - apply eq_of_not_mem_of_mem_insert <| univ₂.symm ▸ mem_univ _ + apply eq_of_mem_insert_of_not_mem <| univ₂.symm ▸ mem_univ _ simpa using mul_not_mem_of_exponent_two h hx hy hxy /-- Any two `IsKleinFour` groups are isomorphic via any equivalence which sends the identity of one diff --git a/Mathlib/Order/CountableDenseLinearOrder.lean b/Mathlib/Order/CountableDenseLinearOrder.lean index 1ab3a1faebbc2..52b1a34e4be76 100644 --- a/Mathlib/Order/CountableDenseLinearOrder.lean +++ b/Mathlib/Order/CountableDenseLinearOrder.lean @@ -82,16 +82,16 @@ lemma exists_orderEmbedding_insert [DenselyOrdered β] [NoMinOrder β] [NoMaxOrd then if hyS : y ∈ S then simpa only [hxS, hyS, ↓reduceDIte, OrderEmbedding.lt_iff_lt, Subtype.mk_lt_mk] else - obtain rfl := Finset.eq_of_not_mem_of_mem_insert hy hyS + obtain rfl := Finset.eq_of_mem_insert_of_not_mem hy hyS simp only [hxS, hyS, ↓reduceDIte] exact hb _ (Finset.mem_image_of_mem _ (Finset.mem_filter.2 ⟨Finset.mem_attach _ _, hxy⟩)) else - obtain rfl := Finset.eq_of_not_mem_of_mem_insert hx hxS + obtain rfl := Finset.eq_of_mem_insert_of_not_mem hx hxS if hyS : y ∈ S then simp only [hxS, hyS, ↓reduceDIte] exact hb' _ (Finset.mem_image_of_mem _ (Finset.mem_filter.2 ⟨Finset.mem_attach _ _, hxy⟩)) - else simp only [Finset.eq_of_not_mem_of_mem_insert hy hyS, lt_self_iff_false] at hxy + else simp only [Finset.eq_of_mem_insert_of_not_mem hy hyS, lt_self_iff_false] at hxy · ext x simp only [Finset.coe_sort_coe, OrderEmbedding.coe_ofStrictMono, Finset.insert_val, Function.comp_apply, Finset.coe_mem, ↓reduceDIte, Subtype.coe_eta] From 495d87ece1fc97b60cedc4b1c9c4e5e9fbb3d7cd Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Fri, 25 Oct 2024 09:55:07 +0000 Subject: [PATCH 410/505] feat: add `cfc_commute_cfc` (#18118) For `cfcHom`, these results are exactly `Commute.map`, but since `cfc` is not actually a homomorphism, it's useful to have this version. --- .../CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean | 5 +++++ .../CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean | 5 +++++ 2 files changed, 10 insertions(+) diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean index e837b34df4812..73105d7530e9b 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean @@ -267,6 +267,11 @@ lemma cfcₙ_cases (P : A → Prop) (a : A) (f : R → R) (h₀ : P 0) · rwa [cfcₙ_apply_of_not_map_zero _ h] · rwa [cfcₙ_apply_of_not_predicate _ h] +lemma cfcₙ_commute_cfcₙ (f g : R → R) (a : A) : Commute (cfcₙ f a) (cfcₙ g a) := by + refine cfcₙ_cases (fun x ↦ Commute x (cfcₙ g a)) a f (by simp) fun hf hf0 ha ↦ ?_ + refine cfcₙ_cases (fun x ↦ Commute _ x) a g (by simp) fun hg hg0 _ ↦ ?_ + exact Commute.all _ _ |>.map _ + variable (R) in include ha in lemma cfcₙ_id : cfcₙ (id : R → R) a = a := diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean index c6070f8408685..b63d226c61d6d 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean @@ -351,6 +351,11 @@ lemma cfc_cases (P : A → Prop) (a : A) (f : R → R) (h₀ : P 0) · rwa [cfc_apply_of_not_predicate _ h] · rwa [cfc_apply_of_not_continuousOn _ h] +lemma cfc_commute_cfc (f g : R → R) (a : A) : Commute (cfc f a) (cfc g a) := by + refine cfc_cases (fun x ↦ Commute x (cfc g a)) a f (by simp) fun hf ha ↦ ?_ + refine cfc_cases (fun x ↦ Commute _ x) a g (by simp) fun hg _ ↦ ?_ + exact Commute.all _ _ |>.map _ + variable (R) in lemma cfc_id (ha : p a := by cfc_tac) : cfc (id : R → R) a = a := cfc_apply (id : R → R) a ▸ cfcHom_id (p := p) ha From 6ca0131ae189339f939abad453cadc81b2811c42 Mon Sep 17 00:00:00 2001 From: Arend Mellendijk Date: Fri, 25 Oct 2024 10:51:48 +0000 Subject: [PATCH 411/505] chore(Tactic/rename_bvar): silence unused tactic linter and add tests (#18139) `rename_bvar` exists only to rename bound variables inside either the goal state or a local hypothesis. This by definition does not change the current goal at all, so the unused tactic linter should be silenced. Add some tests to document this (and the current behaviour). --- Mathlib/Tactic/Linter/UnusedTactic.lean | 1 + Mathlib/Tactic/RenameBVar.lean | 8 +++--- test/renameBvar.lean | 34 +++++++++++++++++++++++++ 3 files changed, 39 insertions(+), 4 deletions(-) create mode 100644 test/renameBvar.lean diff --git a/Mathlib/Tactic/Linter/UnusedTactic.lean b/Mathlib/Tactic/Linter/UnusedTactic.lean index 67eb1b34f7d9f..37f23d5e87ded 100644 --- a/Mathlib/Tactic/Linter/UnusedTactic.lean +++ b/Mathlib/Tactic/Linter/UnusedTactic.lean @@ -93,6 +93,7 @@ initialize allowedRef : IO.Ref (Std.HashSet SyntaxNodeKind) ← |>.insert `change? |>.insert `«tactic#adaptation_note_» |>.insert `tacticSleep_heartbeats_ + |>.insert `Mathlib.Tactic.«tacticRename_bvar_→__» /-- `#allow_unused_tactic` takes an input a space-separated list of identifiers. These identifiers are then allowed by the unused tactic linter: diff --git a/Mathlib/Tactic/RenameBVar.lean b/Mathlib/Tactic/RenameBVar.lean index fe1706f2eacfc..64b015482b45f 100644 --- a/Mathlib/Tactic/RenameBVar.lean +++ b/Mathlib/Tactic/RenameBVar.lean @@ -29,13 +29,13 @@ def renameBVarTarget (mvarId : MVarId) (old new : Name) : MetaM Unit := modifyTarget mvarId fun e ↦ e.renameBVar old new /-- -* `rename_bvar old new` renames all bound variables named `old` to `new` in the target. -* `rename_bvar old new at h` does the same in hypothesis `h`. +* `rename_bvar old → new` renames all bound variables named `old` to `new` in the target. +* `rename_bvar old → new at h` does the same in hypothesis `h`. ```lean example (P : ℕ → ℕ → Prop) (h : ∀ n, ∃ m, P n m) : ∀ l, ∃ m, P l m := by - rename_bvar n q at h -- h is now ∀ (q : ℕ), ∃ (m : ℕ), P q m, - rename_bvar m n -- target is now ∀ (l : ℕ), ∃ (n : ℕ), P k n, + rename_bvar n → q at h -- h is now ∀ (q : ℕ), ∃ (m : ℕ), P q m, + rename_bvar m → n -- target is now ∀ (l : ℕ), ∃ (n : ℕ), P k n, exact h -- Lean does not care about those bound variable names ``` Note: name clashes are resolved automatically. diff --git a/test/renameBvar.lean b/test/renameBvar.lean new file mode 100644 index 0000000000000..07fe22319cebf --- /dev/null +++ b/test/renameBvar.lean @@ -0,0 +1,34 @@ +import Mathlib.Tactic.RenameBVar + +/- This test is somewhat flaky since it depends on the pretty printer. -/ +/-- +info: ℕ : Sort u_1 +P : ℕ → ℕ → Prop +h : ∀ (n : ℕ), ∃ m, P n m +⊢ ∀ (l : ℕ), ∃ m, P l m +--- +info: ℕ : Sort u_1 +P : ℕ → ℕ → Prop +h : ∀ (q : ℕ), ∃ m, P q m +⊢ ∀ (l : ℕ), ∃ m, P l m +--- +info: ℕ : Sort u_1 +P : ℕ → ℕ → Prop +h : ∀ (q : ℕ), ∃ m, P q m +⊢ ∀ (l : ℕ), ∃ n, P l n +--- +info: ℕ : Sort u_1 +P : ℕ → ℕ → Prop +h : ∀ (q : ℕ), ∃ m, P q m +⊢ ∀ (m : ℕ), ∃ n, P m n +-/ +#guard_msgs in +example (P : ℕ → ℕ → Prop) (h : ∀ n, ∃ m, P n m) : ∀ l, ∃ m, P l m := by + trace_state + rename_bvar n → q at h + trace_state + rename_bvar m → n + trace_state + rename_bvar l → m + trace_state + exact h From 086c16c9e1afb5cf1671a508c821c4717fe9050e Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Fri, 25 Oct 2024 11:44:08 +0000 Subject: [PATCH 412/505] chore: split Topology.GDelta (#18215) --- Mathlib.lean | 3 +- .../Constructions/BorelSpace/Basic.lean | 2 +- Mathlib/Topology/Baire/BaireMeasurable.lean | 2 +- Mathlib/Topology/Baire/Lemmas.lean | 2 +- .../{GDelta.lean => GDelta/Basic.lean} | 37 +++---------- Mathlib/Topology/GDelta/UniformSpace.lean | 55 +++++++++++++++++++ Mathlib/Topology/Instances/Irrational.lean | 1 - Mathlib/Topology/Separation/GDelta.lean | 2 +- Mathlib/Topology/UrysohnsLemma.lean | 2 +- 9 files changed, 69 insertions(+), 37 deletions(-) rename Mathlib/Topology/{GDelta.lean => GDelta/Basic.lean} (87%) create mode 100644 Mathlib/Topology/GDelta/UniformSpace.lean diff --git a/Mathlib.lean b/Mathlib.lean index 46bd976d0d3e8..25c3ef177489e 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4779,7 +4779,8 @@ import Mathlib.Topology.FiberBundle.IsHomeomorphicTrivialBundle import Mathlib.Topology.FiberBundle.Trivialization import Mathlib.Topology.FiberPartition import Mathlib.Topology.Filter -import Mathlib.Topology.GDelta +import Mathlib.Topology.GDelta.Basic +import Mathlib.Topology.GDelta.UniformSpace import Mathlib.Topology.Germ import Mathlib.Topology.Gluing import Mathlib.Topology.Hom.ContinuousEval diff --git a/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean b/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean index c5518f43dd87a..f0fe4fcd0af4a 100644 --- a/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean +++ b/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Yury Kudryashov -/ import Mathlib.MeasureTheory.Group.Arithmetic -import Mathlib.Topology.GDelta +import Mathlib.Topology.GDelta.UniformSpace import Mathlib.Topology.Instances.EReal import Mathlib.Topology.Instances.Rat diff --git a/Mathlib/Topology/Baire/BaireMeasurable.lean b/Mathlib/Topology/Baire/BaireMeasurable.lean index 02a160951b95f..1c08bafa7d149 100644 --- a/Mathlib/Topology/Baire/BaireMeasurable.lean +++ b/Mathlib/Topology/Baire/BaireMeasurable.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 Felix Weilacher. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Felix Weilacher -/ -import Mathlib.Topology.GDelta +import Mathlib.Topology.GDelta.UniformSpace import Mathlib.Topology.LocallyClosed import Mathlib.MeasureTheory.Constructions.EventuallyMeasurable import Mathlib.MeasureTheory.Constructions.BorelSpace.Basic diff --git a/Mathlib/Topology/Baire/Lemmas.lean b/Mathlib/Topology/Baire/Lemmas.lean index 7091aebfeb216..926a6354ec0f2 100644 --- a/Mathlib/Topology/Baire/Lemmas.lean +++ b/Mathlib/Topology/Baire/Lemmas.lean @@ -3,7 +3,7 @@ Copyright (c) 2019 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ -import Mathlib.Topology.GDelta +import Mathlib.Topology.GDelta.Basic /-! # Baire spaces diff --git a/Mathlib/Topology/GDelta.lean b/Mathlib/Topology/GDelta/Basic.lean similarity index 87% rename from Mathlib/Topology/GDelta.lean rename to Mathlib/Topology/GDelta/Basic.lean index 94a3d1bec4af6..60e2975ce2246 100644 --- a/Mathlib/Topology/GDelta.lean +++ b/Mathlib/Topology/GDelta/Basic.lean @@ -3,8 +3,9 @@ Copyright (c) 2019 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel, Yury Kudryashov -/ -import Mathlib.Topology.UniformSpace.Basic +import Mathlib.Algebra.Group.Defs import Mathlib.Order.Filter.CountableInter +import Mathlib.Topology.Basic /-! # `Gδ` sets @@ -24,8 +25,7 @@ In this file we define `Gδ` sets and prove their basic properties. ## Main results -We prove that finite or countable intersections of Gδ sets are Gδ sets. We also prove that the -continuity set of a function from a topological space to an (e)metric space is a Gδ set. +We prove that finite or countable intersections of Gδ sets are Gδ sets. - `isClosed_isNowhereDense_iff_compl`: a closed set is nowhere dense iff its complement is open and dense @@ -33,16 +33,19 @@ its complement is open and dense union of nowhere dense sets - subsets of meagre sets are meagre; countable unions of meagre sets are meagre +See `Mathlib.Topology.GDelta.UniformSpace` for the proof that +continuity set of a function from a topological space to a uniform space is a Gδ set. + ## Tags Gδ set, residual set, nowhere dense set, meagre set -/ +assert_not_exists UniformSpace noncomputable section open Topology TopologicalSpace Filter Encodable Set -open scoped Uniformity variable {X Y ι : Type*} {ι' : Sort*} @@ -154,34 +157,8 @@ alias isGδ_biUnion := IsGδ.biUnion theorem IsGδ.iUnion [Finite ι'] {f : ι' → Set X} (h : ∀ i, IsGδ (f i)) : IsGδ (⋃ i, f i) := .sUnion (finite_range _) <| forall_mem_range.2 h -theorem IsClosed.isGδ {X : Type*} [UniformSpace X] [IsCountablyGenerated (𝓤 X)] {s : Set X} - (hs : IsClosed s) : IsGδ s := by - rcases (@uniformity_hasBasis_open X _).exists_antitone_subbasis with ⟨U, hUo, hU, -⟩ - rw [← hs.closure_eq, ← hU.biInter_biUnion_ball] - refine .biInter (to_countable _) fun n _ => IsOpen.isGδ ?_ - exact isOpen_biUnion fun x _ => UniformSpace.isOpen_ball _ (hUo _).2 - end IsGδ -section ContinuousAt - -variable [TopologicalSpace X] - -/-- The set of points where a function is continuous is a Gδ set. -/ -theorem IsGδ.setOf_continuousAt [UniformSpace Y] [IsCountablyGenerated (𝓤 Y)] (f : X → Y) : - IsGδ { x | ContinuousAt f x } := by - obtain ⟨U, _, hU⟩ := (@uniformity_hasBasis_open_symmetric Y _).exists_antitone_subbasis - simp only [Uniform.continuousAt_iff_prod, nhds_prod_eq] - simp only [(nhds_basis_opens _).prod_self.tendsto_iff hU.toHasBasis, forall_prop_of_true, - setOf_forall, id] - refine .iInter fun k ↦ IsOpen.isGδ <| isOpen_iff_mem_nhds.2 fun x ↦ ?_ - rintro ⟨s, ⟨hsx, hso⟩, hsU⟩ - filter_upwards [IsOpen.mem_nhds hso hsx] with _ hy using ⟨s, ⟨hy, hso⟩, hsU⟩ - -@[deprecated (since := "2024-02-15")] alias isGδ_setOf_continuousAt := IsGδ.setOf_continuousAt - -end ContinuousAt - section residual variable [TopologicalSpace X] diff --git a/Mathlib/Topology/GDelta/UniformSpace.lean b/Mathlib/Topology/GDelta/UniformSpace.lean new file mode 100644 index 0000000000000..aa957353a7f05 --- /dev/null +++ b/Mathlib/Topology/GDelta/UniformSpace.lean @@ -0,0 +1,55 @@ +/- +Copyright (c) 2019 Sébastien Gouëzel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sébastien Gouëzel, Yury Kudryashov +-/ +import Mathlib.Topology.GDelta.Basic +import Mathlib.Topology.UniformSpace.Basic +import Mathlib.Order.Filter.CountableInter + +/-! +# `Gδ` sets and uniform spaces + +## Main results +We prove that the continuity set of a function from a topological space to a uniform space is a +Gδ set. + +-/ + + +noncomputable section + +open Topology TopologicalSpace Filter Encodable Set +open scoped Uniformity + +variable {X Y ι : Type*} {ι' : Sort*} + +section IsGδ + +theorem IsClosed.isGδ {X : Type*} [UniformSpace X] [IsCountablyGenerated (𝓤 X)] {s : Set X} + (hs : IsClosed s) : IsGδ s := by + rcases (@uniformity_hasBasis_open X _).exists_antitone_subbasis with ⟨U, hUo, hU, -⟩ + rw [← hs.closure_eq, ← hU.biInter_biUnion_ball] + refine .biInter (to_countable _) fun n _ => IsOpen.isGδ ?_ + exact isOpen_biUnion fun x _ => UniformSpace.isOpen_ball _ (hUo _).2 + +end IsGδ + +section ContinuousAt + +variable [TopologicalSpace X] + +/-- The set of points where a function is continuous is a Gδ set. -/ +theorem IsGδ.setOf_continuousAt [UniformSpace Y] [IsCountablyGenerated (𝓤 Y)] (f : X → Y) : + IsGδ { x | ContinuousAt f x } := by + obtain ⟨U, _, hU⟩ := (@uniformity_hasBasis_open_symmetric Y _).exists_antitone_subbasis + simp only [Uniform.continuousAt_iff_prod, nhds_prod_eq] + simp only [(nhds_basis_opens _).prod_self.tendsto_iff hU.toHasBasis, forall_prop_of_true, + setOf_forall, id] + refine .iInter fun k ↦ IsOpen.isGδ <| isOpen_iff_mem_nhds.2 fun x ↦ ?_ + rintro ⟨s, ⟨hsx, hso⟩, hsU⟩ + filter_upwards [IsOpen.mem_nhds hso hsx] with _ hy using ⟨s, ⟨hy, hso⟩, hsU⟩ + +@[deprecated (since := "2024-02-15")] alias isGδ_setOf_continuousAt := IsGδ.setOf_continuousAt + +end ContinuousAt diff --git a/Mathlib/Topology/Instances/Irrational.lean b/Mathlib/Topology/Instances/Irrational.lean index ad7ac6309e9fa..f395be578fdff 100644 --- a/Mathlib/Topology/Instances/Irrational.lean +++ b/Mathlib/Topology/Instances/Irrational.lean @@ -5,7 +5,6 @@ Authors: Yury Kudryashov -/ import Mathlib.Data.Real.Irrational import Mathlib.Data.Rat.Encodable -import Mathlib.Topology.GDelta import Mathlib.Topology.Separation.GDelta /-! diff --git a/Mathlib/Topology/Separation/GDelta.lean b/Mathlib/Topology/Separation/GDelta.lean index 3b58de08a18ee..e7cf8efc00758 100644 --- a/Mathlib/Topology/Separation/GDelta.lean +++ b/Mathlib/Topology/Separation/GDelta.lean @@ -7,8 +7,8 @@ import Mathlib.Topology.Compactness.Lindelof import Mathlib.Topology.Compactness.SigmaCompact import Mathlib.Topology.Connected.TotallyDisconnected import Mathlib.Topology.Inseparable -import Mathlib.Topology.GDelta import Mathlib.Topology.Separation.Basic +import Mathlib.Topology.GDelta.Basic /-! # Separation properties of topological spaces. diff --git a/Mathlib/Topology/UrysohnsLemma.lean b/Mathlib/Topology/UrysohnsLemma.lean index 705eb5ca08358..f645df48176ff 100644 --- a/Mathlib/Topology/UrysohnsLemma.lean +++ b/Mathlib/Topology/UrysohnsLemma.lean @@ -6,9 +6,9 @@ Authors: Yury Kudryashov import Mathlib.Analysis.Normed.Affine.AddTorsor import Mathlib.LinearAlgebra.AffineSpace.Ordered import Mathlib.Topology.ContinuousMap.Basic -import Mathlib.Topology.GDelta import Mathlib.Analysis.NormedSpace.FunctionSeries import Mathlib.Analysis.SpecificLimits.Basic +import Mathlib.Topology.GDelta.Basic /-! # Urysohn's lemma From 21128da4eb35f0ab7bb8165a856f11d38de987b3 Mon Sep 17 00:00:00 2001 From: Markus Himmel Date: Fri, 25 Oct 2024 12:15:34 +0000 Subject: [PATCH 413/505] chore: split StructuredArrow.lean to reduce imports (#18227) Found using the new `lake exe unused`. Co-authored-by: Markus Himmel --- Mathlib.lean | 3 +- .../Adjunction/AdjointFunctorTheorems.lean | 1 + Mathlib/CategoryTheory/Adjunction/Comma.lean | 2 +- .../CategoryTheory/Bicategory/Extension.lean | 2 +- Mathlib/CategoryTheory/Comma/Over.lean | 2 +- .../Basic.lean} | 16 ------ .../Comma/StructuredArrow/Small.lean | 50 +++++++++++++++++++ Mathlib/CategoryTheory/Elements.lean | 2 +- .../Functor/KanExtension/Basic.lean | 2 +- Mathlib/CategoryTheory/Limits/Comma.lean | 2 +- Mathlib/CategoryTheory/Limits/Final.lean | 2 +- .../CategoryTheory/Limits/FinallySmall.lean | 1 + .../Localization/StructuredArrow.lean | 2 +- Mathlib/CategoryTheory/Sites/Sieves.lean | 1 + 14 files changed, 63 insertions(+), 25 deletions(-) rename Mathlib/CategoryTheory/Comma/{StructuredArrow.lean => StructuredArrow/Basic.lean} (97%) create mode 100644 Mathlib/CategoryTheory/Comma/StructuredArrow/Small.lean diff --git a/Mathlib.lean b/Mathlib.lean index 25c3ef177489e..a58a3a3939b0d 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1541,7 +1541,8 @@ import Mathlib.CategoryTheory.Comma.Basic import Mathlib.CategoryTheory.Comma.Over import Mathlib.CategoryTheory.Comma.Presheaf.Basic import Mathlib.CategoryTheory.Comma.Presheaf.Colimit -import Mathlib.CategoryTheory.Comma.StructuredArrow +import Mathlib.CategoryTheory.Comma.StructuredArrow.Basic +import Mathlib.CategoryTheory.Comma.StructuredArrow.Small import Mathlib.CategoryTheory.ComposableArrows import Mathlib.CategoryTheory.ConcreteCategory.Basic import Mathlib.CategoryTheory.ConcreteCategory.Bundled diff --git a/Mathlib/CategoryTheory/Adjunction/AdjointFunctorTheorems.lean b/Mathlib/CategoryTheory/Adjunction/AdjointFunctorTheorems.lean index 3d5e083b0ab9b..6d42354bbaaa3 100644 --- a/Mathlib/CategoryTheory/Adjunction/AdjointFunctorTheorems.lean +++ b/Mathlib/CategoryTheory/Adjunction/AdjointFunctorTheorems.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Bhavik Mehta. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Bhavik Mehta -/ +import Mathlib.CategoryTheory.Comma.StructuredArrow.Small import Mathlib.CategoryTheory.Generator import Mathlib.CategoryTheory.Limits.ConeCategory import Mathlib.CategoryTheory.Limits.Constructions.WeaklyInitial diff --git a/Mathlib/CategoryTheory/Adjunction/Comma.lean b/Mathlib/CategoryTheory/Adjunction/Comma.lean index 3b59a2dd44bdc..6f4474b55e705 100644 --- a/Mathlib/CategoryTheory/Adjunction/Comma.lean +++ b/Mathlib/CategoryTheory/Adjunction/Comma.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Bhavik Mehta -/ import Mathlib.CategoryTheory.Adjunction.Basic -import Mathlib.CategoryTheory.Comma.StructuredArrow +import Mathlib.CategoryTheory.Comma.StructuredArrow.Basic import Mathlib.CategoryTheory.PUnit /-! diff --git a/Mathlib/CategoryTheory/Bicategory/Extension.lean b/Mathlib/CategoryTheory/Bicategory/Extension.lean index 898bc3c43dd3e..5f371e5ffbfee 100644 --- a/Mathlib/CategoryTheory/Bicategory/Extension.lean +++ b/Mathlib/CategoryTheory/Bicategory/Extension.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yuma Mizuno -/ import Mathlib.CategoryTheory.Bicategory.Basic -import Mathlib.CategoryTheory.Comma.StructuredArrow +import Mathlib.CategoryTheory.Comma.StructuredArrow.Basic /-! # Extensions and lifts in bicategories diff --git a/Mathlib/CategoryTheory/Comma/Over.lean b/Mathlib/CategoryTheory/Comma/Over.lean index 3b9a32d2c8eee..2d0ed2bd61b9c 100644 --- a/Mathlib/CategoryTheory/Comma/Over.lean +++ b/Mathlib/CategoryTheory/Comma/Over.lean @@ -3,7 +3,7 @@ Copyright (c) 2019 Johan Commelin. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin, Bhavik Mehta -/ -import Mathlib.CategoryTheory.Comma.StructuredArrow +import Mathlib.CategoryTheory.Comma.StructuredArrow.Basic import Mathlib.CategoryTheory.PUnit import Mathlib.CategoryTheory.Functor.ReflectsIso import Mathlib.CategoryTheory.Functor.EpiMono diff --git a/Mathlib/CategoryTheory/Comma/StructuredArrow.lean b/Mathlib/CategoryTheory/Comma/StructuredArrow/Basic.lean similarity index 97% rename from Mathlib/CategoryTheory/Comma/StructuredArrow.lean rename to Mathlib/CategoryTheory/Comma/StructuredArrow/Basic.lean index d86b41669eacf..c0d1729054be6 100644 --- a/Mathlib/CategoryTheory/Comma/StructuredArrow.lean +++ b/Mathlib/CategoryTheory/Comma/StructuredArrow/Basic.lean @@ -6,8 +6,6 @@ Authors: Adam Topaz, Kim Morrison import Mathlib.CategoryTheory.Comma.Basic import Mathlib.CategoryTheory.PUnit import Mathlib.CategoryTheory.Limits.Shapes.Terminal -import Mathlib.CategoryTheory.EssentiallySmall -import Mathlib.Logic.Small.Set /-! # The category of "structured arrows" @@ -331,13 +329,6 @@ noncomputable instance isEquivalenceMap₂ end -instance small_proj_preimage_of_locallySmall {𝒢 : Set C} [Small.{v₁} 𝒢] [LocallySmall.{v₁} D] : - Small.{v₁} ((proj S T).obj ⁻¹' 𝒢) := by - suffices (proj S T).obj ⁻¹' 𝒢 = Set.range fun f : ΣG : 𝒢, S ⟶ T.obj G => mk f.2 by - rw [this] - infer_instance - exact Set.ext fun X => ⟨fun h => ⟨⟨⟨_, h⟩, X.hom⟩, (eq_mk _).symm⟩, by aesop_cat⟩ - /-- A structured arrow is called universal if it is initial. -/ abbrev IsUniversal (f : StructuredArrow S T) := IsInitial f @@ -674,13 +665,6 @@ noncomputable instance isEquivalenceMap₂ end -instance small_proj_preimage_of_locallySmall {𝒢 : Set C} [Small.{v₁} 𝒢] [LocallySmall.{v₁} D] : - Small.{v₁} ((proj S T).obj ⁻¹' 𝒢) := by - suffices (proj S T).obj ⁻¹' 𝒢 = Set.range fun f : ΣG : 𝒢, S.obj G ⟶ T => mk f.2 by - rw [this] - infer_instance - exact Set.ext fun X => ⟨fun h => ⟨⟨⟨_, h⟩, X.hom⟩, (eq_mk _).symm⟩, by aesop_cat⟩ - /-- A costructured arrow is called universal if it is terminal. -/ abbrev IsUniversal (f : CostructuredArrow S T) := IsTerminal f diff --git a/Mathlib/CategoryTheory/Comma/StructuredArrow/Small.lean b/Mathlib/CategoryTheory/Comma/StructuredArrow/Small.lean new file mode 100644 index 0000000000000..a31121a52e88a --- /dev/null +++ b/Mathlib/CategoryTheory/Comma/StructuredArrow/Small.lean @@ -0,0 +1,50 @@ +/- +Copyright (c) 2022 Markus Himmel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Markus Himmel +-/ +import Mathlib.CategoryTheory.Comma.StructuredArrow.Basic +import Mathlib.CategoryTheory.EssentiallySmall +import Mathlib.Logic.Small.Set + +/-! +# Small sets in the category of structured arrows + +Here we prove a technical result about small sets in the category of structured arrows that will +be used in the proof of the Special Adjoint Functor Theorem. +-/ + +namespace CategoryTheory + +-- morphism levels before object levels. See note [CategoryTheory universes]. +universe v₁ v₂ u₁ u₂ + +variable {C : Type u₁} [Category.{v₁} C] {D : Type u₂} [Category.{v₂} D] + +namespace StructuredArrow + +variable {S : D} {T : C ⥤ D} + +instance small_proj_preimage_of_locallySmall {𝒢 : Set C} [Small.{v₁} 𝒢] [LocallySmall.{v₁} D] : + Small.{v₁} ((proj S T).obj ⁻¹' 𝒢) := by + suffices (proj S T).obj ⁻¹' 𝒢 = Set.range fun f : ΣG : 𝒢, S ⟶ T.obj G => mk f.2 by + rw [this] + infer_instance + exact Set.ext fun X => ⟨fun h => ⟨⟨⟨_, h⟩, X.hom⟩, (eq_mk _).symm⟩, by aesop_cat⟩ + +end StructuredArrow + +namespace CostructuredArrow + +variable {S : C ⥤ D} {T : D} + +instance small_proj_preimage_of_locallySmall {𝒢 : Set C} [Small.{v₁} 𝒢] [LocallySmall.{v₁} D] : + Small.{v₁} ((proj S T).obj ⁻¹' 𝒢) := by + suffices (proj S T).obj ⁻¹' 𝒢 = Set.range fun f : ΣG : 𝒢, S.obj G ⟶ T => mk f.2 by + rw [this] + infer_instance + exact Set.ext fun X => ⟨fun h => ⟨⟨⟨_, h⟩, X.hom⟩, (eq_mk _).symm⟩, by aesop_cat⟩ + +end CostructuredArrow + +end CategoryTheory diff --git a/Mathlib/CategoryTheory/Elements.lean b/Mathlib/CategoryTheory/Elements.lean index cf261e1d8d7fc..2c682ecb881f9 100644 --- a/Mathlib/CategoryTheory/Elements.lean +++ b/Mathlib/CategoryTheory/Elements.lean @@ -3,7 +3,7 @@ Copyright (c) 2019 Kim Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison -/ -import Mathlib.CategoryTheory.Comma.StructuredArrow +import Mathlib.CategoryTheory.Comma.StructuredArrow.Basic import Mathlib.CategoryTheory.Groupoid import Mathlib.CategoryTheory.PUnit diff --git a/Mathlib/CategoryTheory/Functor/KanExtension/Basic.lean b/Mathlib/CategoryTheory/Functor/KanExtension/Basic.lean index 6b966a9ee45ed..7c08315b07111 100644 --- a/Mathlib/CategoryTheory/Functor/KanExtension/Basic.lean +++ b/Mathlib/CategoryTheory/Functor/KanExtension/Basic.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 Joël Riou. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Joël Riou -/ -import Mathlib.CategoryTheory.Comma.StructuredArrow +import Mathlib.CategoryTheory.Comma.StructuredArrow.Basic import Mathlib.CategoryTheory.Limits.Shapes.Equivalence import Mathlib.CategoryTheory.Limits.Preserves.Shapes.Terminal diff --git a/Mathlib/CategoryTheory/Limits/Comma.lean b/Mathlib/CategoryTheory/Limits/Comma.lean index a54e3c6a62630..5e5e2804c305b 100644 --- a/Mathlib/CategoryTheory/Limits/Comma.lean +++ b/Mathlib/CategoryTheory/Limits/Comma.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Bhavik Mehta -/ import Mathlib.CategoryTheory.Comma.Arrow -import Mathlib.CategoryTheory.Comma.StructuredArrow +import Mathlib.CategoryTheory.Comma.StructuredArrow.Basic import Mathlib.CategoryTheory.Limits.Constructions.EpiMono import Mathlib.CategoryTheory.Limits.Creates import Mathlib.CategoryTheory.Limits.Unit diff --git a/Mathlib/CategoryTheory/Limits/Final.lean b/Mathlib/CategoryTheory/Limits/Final.lean index 12c39f6153dc8..8dea5b13733db 100644 --- a/Mathlib/CategoryTheory/Limits/Final.lean +++ b/Mathlib/CategoryTheory/Limits/Final.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Kim Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison -/ -import Mathlib.CategoryTheory.Comma.StructuredArrow +import Mathlib.CategoryTheory.Comma.StructuredArrow.Basic import Mathlib.CategoryTheory.IsConnected import Mathlib.CategoryTheory.Limits.Preserves.Shapes.Terminal import Mathlib.CategoryTheory.Limits.Shapes.Types diff --git a/Mathlib/CategoryTheory/Limits/FinallySmall.lean b/Mathlib/CategoryTheory/Limits/FinallySmall.lean index 90703593a078a..3613b80d171c2 100644 --- a/Mathlib/CategoryTheory/Limits/FinallySmall.lean +++ b/Mathlib/CategoryTheory/Limits/FinallySmall.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Markus Himmel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Markus Himmel -/ +import Mathlib.Logic.Small.Set import Mathlib.CategoryTheory.Filtered.Final /-! diff --git a/Mathlib/CategoryTheory/Localization/StructuredArrow.lean b/Mathlib/CategoryTheory/Localization/StructuredArrow.lean index 91c9d508c56f6..96c0af25daccb 100644 --- a/Mathlib/CategoryTheory/Localization/StructuredArrow.lean +++ b/Mathlib/CategoryTheory/Localization/StructuredArrow.lean @@ -5,7 +5,7 @@ Authors: Joël Riou -/ import Mathlib.CategoryTheory.Localization.HomEquiv import Mathlib.CategoryTheory.Localization.Opposite -import Mathlib.CategoryTheory.Comma.StructuredArrow +import Mathlib.CategoryTheory.Comma.StructuredArrow.Basic /-! # Induction principles for structured and costructured arrows diff --git a/Mathlib/CategoryTheory/Sites/Sieves.lean b/Mathlib/CategoryTheory/Sites/Sieves.lean index d974d994459f0..66591c5a8f117 100644 --- a/Mathlib/CategoryTheory/Sites/Sieves.lean +++ b/Mathlib/CategoryTheory/Sites/Sieves.lean @@ -3,6 +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, Edward Ayers -/ +import Mathlib.Data.Set.Lattice import Mathlib.CategoryTheory.Limits.Shapes.Pullback.HasPullback /-! From d09d074b8b2b5d8f6a89f259cd8751d5a2a45c9e Mon Sep 17 00:00:00 2001 From: Xavier Roblot Date: Fri, 25 Oct 2024 12:49:04 +0000 Subject: [PATCH 414/505] chore(NumberField/Embeddings): change `Nr(Real|Complex)Places` to `nr(Real|Complex)Places` (#18140) As noted by @acmepjz, `NrRealPlaces` and `NrComplexPlaces` should be called `nrRealPlaces` and `nrComplexPlaces`. This is fixed in this PR. Co-authored-by: Xavier Roblot <46200072+xroblot@users.noreply.github.com> --- .../NumberTheory/Cyclotomic/Embeddings.lean | 4 +-- .../NumberField/CanonicalEmbedding/Basic.lean | 8 +++--- .../CanonicalEmbedding/ConvexBody.lean | 26 +++++++++---------- .../NumberTheory/NumberField/ClassNumber.lean | 4 +-- .../NumberField/Discriminant/Basic.lean | 22 ++++++++-------- .../NumberTheory/NumberField/Embeddings.lean | 22 +++++++++------- 6 files changed, 45 insertions(+), 41 deletions(-) diff --git a/Mathlib/NumberTheory/Cyclotomic/Embeddings.lean b/Mathlib/NumberTheory/Cyclotomic/Embeddings.lean index 642d6933fad59..a6b2d1fc27387 100644 --- a/Mathlib/NumberTheory/Cyclotomic/Embeddings.lean +++ b/Mathlib/NumberTheory/Cyclotomic/Embeddings.lean @@ -30,7 +30,7 @@ of `K`. -/ theorem nrRealPlaces_eq_zero [IsCyclotomicExtension {n} ℚ K] (hn : 2 < n) : haveI := IsCyclotomicExtension.numberField {n} ℚ K - NrRealPlaces K = 0 := by + nrRealPlaces K = 0 := by have := IsCyclotomicExtension.numberField {n} ℚ K apply (IsCyclotomicExtension.zeta_spec n ℚ K).nrRealPlaces_eq_zero_of_two_lt hn @@ -40,7 +40,7 @@ variable (n) of `K`. Note that this uses `1 / 2 = 0` in the cases `n = 1, 2`. -/ theorem nrComplexPlaces_eq_totient_div_two [h : IsCyclotomicExtension {n} ℚ K] : haveI := IsCyclotomicExtension.numberField {n} ℚ K - NrComplexPlaces K = φ n / 2 := by + nrComplexPlaces K = φ n / 2 := by have := IsCyclotomicExtension.numberField {n} ℚ K by_cases hn : 2 < n · obtain ⟨k, hk : φ n = k + k⟩ := totient_even hn diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean index f119b3f3072eb..0ff7cd0991a72 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean @@ -208,7 +208,7 @@ instance [NumberField K] : Nontrivial (mixedSpace K) := by protected theorem finrank [NumberField K] : finrank ℝ (mixedSpace K) = finrank ℚ K := by classical rw [finrank_prod, finrank_pi, finrank_pi_fintype, Complex.finrank_real_complex, sum_const, - card_univ, ← NrRealPlaces, ← NrComplexPlaces, ← card_real_embeddings, Algebra.id.smul_eq_mul, + card_univ, ← nrRealPlaces, ← nrComplexPlaces, ← card_real_embeddings, Algebra.id.smul_eq_mul, mul_comm, ← card_complex_embeddings, ← NumberField.Embeddings.card K ℂ, Fintype.card_subtype_compl, Nat.add_sub_of_le (Fintype.card_subtype_le _)] @@ -517,7 +517,7 @@ def matrixToStdBasis : Matrix (index K) (index K) ℂ := (blockDiagonal (fun _ => (2 : ℂ)⁻¹ • !![1, 1; - I, I])) theorem det_matrixToStdBasis : - (matrixToStdBasis K).det = (2⁻¹ * I) ^ NrComplexPlaces K := + (matrixToStdBasis K).det = (2⁻¹ * I) ^ nrComplexPlaces K := calc _ = ∏ _k : { w : InfinitePlace K // IsComplex w }, det ((2 : ℂ)⁻¹ • !![1, 1; -I, I]) := by rw [matrixToStdBasis, det_fromBlocks_zero₂₁, det_diagonal, prod_const_one, one_mul, @@ -590,8 +590,8 @@ def latticeBasis : -- and it's a basis since it has the right cardinality refine basisOfLinearIndependentOfCardEqFinrank this ?_ rw [← finrank_eq_card_chooseBasisIndex, RingOfIntegers.rank, finrank_prod, finrank_pi, - finrank_pi_fintype, Complex.finrank_real_complex, sum_const, card_univ, ← NrRealPlaces, - ← NrComplexPlaces, ← card_real_embeddings, Algebra.id.smul_eq_mul, mul_comm, + finrank_pi_fintype, Complex.finrank_real_complex, sum_const, card_univ, ← nrRealPlaces, + ← nrComplexPlaces, ← card_real_embeddings, Algebra.id.smul_eq_mul, mul_comm, ← card_complex_embeddings, ← NumberField.Embeddings.card K ℂ, Fintype.card_subtype_compl, Nat.add_sub_of_le (Fintype.card_subtype_le _)] diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/ConvexBody.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/ConvexBody.lean index 36a06aabe738b..561216bee288f 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/ConvexBody.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/ConvexBody.lean @@ -88,7 +88,7 @@ instance : NoAtoms (volume : Measure (mixedSpace K)) := by /-- The fudge factor that appears in the formula for the volume of `convexBodyLT`. -/ noncomputable abbrev convexBodyLTFactor : ℝ≥0 := - (2 : ℝ≥0) ^ NrRealPlaces K * NNReal.pi ^ NrComplexPlaces K + (2 : ℝ≥0) ^ nrRealPlaces K * NNReal.pi ^ nrComplexPlaces K theorem convexBodyLTFactor_ne_zero : convexBodyLTFactor K ≠ 0 := mul_ne_zero (pow_ne_zero _ two_ne_zero) (pow_ne_zero _ pi_ne_zero) @@ -104,10 +104,10 @@ theorem convexBodyLT_volume : _ = (∏ x : {w // InfinitePlace.IsReal w}, ENNReal.ofReal (2 * (f x.val))) * ∏ x : {w // InfinitePlace.IsComplex w}, ENNReal.ofReal (f x.val) ^ 2 * NNReal.pi := by simp_rw [volume_eq_prod, prod_prod, volume_pi, pi_pi, Real.volume_ball, Complex.volume_ball] - _ = ((2 : ℝ≥0) ^ NrRealPlaces K + _ = ((2 : ℝ≥0) ^ nrRealPlaces K * (∏ x : {w // InfinitePlace.IsReal w}, ENNReal.ofReal (f x.val))) * ((∏ x : {w // IsComplex w}, ENNReal.ofReal (f x.val) ^ 2) * - NNReal.pi ^ NrComplexPlaces K) := by + NNReal.pi ^ nrComplexPlaces K) := by simp_rw [ofReal_mul (by norm_num : 0 ≤ (2 : ℝ)), Finset.prod_mul_distrib, Finset.prod_const, Finset.card_univ, ofReal_ofNat, ofReal_coe_nnreal, coe_ofNat] _ = (convexBodyLTFactor K) * ((∏ x : {w // InfinitePlace.IsReal w}, .ofReal (f x.val)) * @@ -205,7 +205,7 @@ variable [NumberField K] /-- The fudge factor that appears in the formula for the volume of `convexBodyLT'`. -/ noncomputable abbrev convexBodyLT'Factor : ℝ≥0 := - (2 : ℝ≥0) ^ (NrRealPlaces K + 2) * NNReal.pi ^ (NrComplexPlaces K - 1) + (2 : ℝ≥0) ^ (nrRealPlaces K + 2) * NNReal.pi ^ (nrComplexPlaces K - 1) theorem convexBodyLT'Factor_ne_zero : convexBodyLT'Factor K ≠ 0 := mul_ne_zero (pow_ne_zero _ two_ne_zero) (pow_ne_zero _ pi_ne_zero) @@ -238,10 +238,10 @@ theorem convexBodyLT'_volume : · refine Finset.prod_congr rfl (fun w' hw' ↦ ?_) rw [if_neg (Finset.ne_of_mem_erase hw'), Complex.volume_ball] · simpa only [ite_true] using vol_box (f w₀) - _ = ((2 : ℝ≥0) ^ NrRealPlaces K * + _ = ((2 : ℝ≥0) ^ nrRealPlaces K * (∏ x : {w // InfinitePlace.IsReal w}, ENNReal.ofReal (f x.val))) * ((∏ x ∈ Finset.univ.erase w₀, ENNReal.ofReal (f x.val) ^ 2) * - ↑pi ^ (NrComplexPlaces K - 1) * (4 * (f w₀) ^ 2)) := by + ↑pi ^ (nrComplexPlaces K - 1) * (4 * (f w₀) ^ 2)) := by simp_rw [ofReal_mul (by norm_num : 0 ≤ (2 : ℝ)), Finset.prod_mul_distrib, Finset.prod_const, Finset.card_erase_of_mem (Finset.mem_univ _), Finset.card_univ, ofReal_ofNat, ofReal_coe_nnreal, coe_ofNat] @@ -385,7 +385,7 @@ theorem convexBodySum_compact : IsCompact (convexBodySum K B) := by /-- The fudge factor that appears in the formula for the volume of `convexBodyLt`. -/ noncomputable abbrev convexBodySumFactor : ℝ≥0 := - (2 : ℝ≥0) ^ NrRealPlaces K * (NNReal.pi / 2) ^ NrComplexPlaces K / (finrank ℚ K).factorial + (2 : ℝ≥0) ^ nrRealPlaces K * (NNReal.pi / 2) ^ nrComplexPlaces K / (finrank ℚ K).factorial theorem convexBodySumFactor_ne_zero : convexBodySumFactor K ≠ 0 := by refine div_ne_zero ?_ <| Nat.cast_ne_zero.mpr (Nat.factorial_ne_zero _) @@ -418,7 +418,7 @@ theorem convexBodySum_volume : simp_rw [mixedEmbedding.finrank, div_one, Gamma_nat_eq_factorial, ofReal_div_of_pos (Nat.cast_pos.mpr (Nat.factorial_pos _)), Real.rpow_one, ofReal_natCast] suffices ∫ x : mixedSpace K, exp (-convexBodySumFun x) = - (2 : ℝ) ^ NrRealPlaces K * (π / 2) ^ NrComplexPlaces K by + (2 : ℝ) ^ nrRealPlaces K * (π / 2) ^ nrComplexPlaces K by rw [this, convexBodySumFactor, ofReal_mul (by positivity), ofReal_pow zero_le_two, ofReal_pow (by positivity), ofReal_div_of_pos zero_lt_two, ofReal_ofNat, ← NNReal.coe_real_pi, ofReal_coe_nnreal, coe_div (Nat.cast_ne_zero.mpr @@ -429,17 +429,17 @@ theorem convexBodySum_volume : (∫ x : {w : InfinitePlace K // IsComplex w} → ℂ, ∏ w, exp (- 2 * ‖x w‖)) := by simp_rw [convexBodySumFun_apply', neg_add, ← neg_mul, Finset.mul_sum, ← Finset.sum_neg_distrib, exp_add, exp_sum, ← integral_prod_mul, volume_eq_prod] - _ = (∫ x : ℝ, exp (-|x|)) ^ NrRealPlaces K * - (∫ x : ℂ, Real.exp (-2 * ‖x‖)) ^ NrComplexPlaces K := by + _ = (∫ x : ℝ, exp (-|x|)) ^ nrRealPlaces K * + (∫ x : ℂ, Real.exp (-2 * ‖x‖)) ^ nrComplexPlaces K := by rw [integral_fintype_prod_eq_pow _ (fun x => exp (- ‖x‖)), integral_fintype_prod_eq_pow _ (fun x => exp (- 2 * ‖x‖))] simp_rw [norm_eq_abs] - _ = (2 * Gamma (1 / 1 + 1)) ^ NrRealPlaces K * - (π * (2 : ℝ) ^ (-(2 : ℝ) / 1) * Gamma (2 / 1 + 1)) ^ NrComplexPlaces K := by + _ = (2 * Gamma (1 / 1 + 1)) ^ nrRealPlaces K * + (π * (2 : ℝ) ^ (-(2 : ℝ) / 1) * Gamma (2 / 1 + 1)) ^ nrComplexPlaces K := by rw [integral_comp_abs (f := fun x => exp (- x)), ← integral_exp_neg_rpow zero_lt_one, ← Complex.integral_exp_neg_mul_rpow le_rfl zero_lt_two] simp_rw [Real.rpow_one] - _ = (2 : ℝ) ^ NrRealPlaces K * (π / 2) ^ NrComplexPlaces K := by + _ = (2 : ℝ) ^ nrRealPlaces K * (π / 2) ^ nrComplexPlaces K := by simp_rw [div_one, one_add_one_eq_two, Gamma_add_one two_ne_zero, Gamma_two, mul_one, mul_assoc, ← Real.rpow_add_one two_ne_zero, show (-2 : ℝ) + 1 = -1 by norm_num, Real.rpow_neg_one] diff --git a/Mathlib/NumberTheory/NumberField/ClassNumber.lean b/Mathlib/NumberTheory/NumberField/ClassNumber.lean index 769d1f2df3687..0c87862166478 100644 --- a/Mathlib/NumberTheory/NumberField/ClassNumber.lean +++ b/Mathlib/NumberTheory/NumberField/ClassNumber.lean @@ -46,7 +46,7 @@ open scoped nonZeroDivisors Real theorem exists_ideal_in_class_of_norm_le (C : ClassGroup (𝓞 K)) : ∃ I : (Ideal (𝓞 K))⁰, ClassGroup.mk0 I = C ∧ - Ideal.absNorm (I : Ideal (𝓞 K)) ≤ (4 / π) ^ NrComplexPlaces K * + Ideal.absNorm (I : Ideal (𝓞 K)) ≤ (4 / π) ^ nrComplexPlaces K * ((finrank ℚ K).factorial / (finrank ℚ K) ^ (finrank ℚ K) * Real.sqrt |discr K|) := by obtain ⟨J, hJ⟩ := ClassGroup.mk0_surjective C⁻¹ obtain ⟨_, ⟨a, ha, rfl⟩, h_nz, h_nm⟩ := @@ -70,7 +70,7 @@ theorem exists_ideal_in_class_of_norm_le (C : ClassGroup (𝓞 K)) : exact Nat.cast_pos.mpr <| Nat.pos_of_ne_zero <| Ideal.absNorm_ne_zero_of_nonZeroDivisors J theorem _root_.RingOfIntegers.isPrincipalIdealRing_of_abs_discr_lt - (h : |discr K| < (2 * (π / 4) ^ NrComplexPlaces K * + (h : |discr K| < (2 * (π / 4) ^ nrComplexPlaces K * ((finrank ℚ K) ^ (finrank ℚ K) / (finrank ℚ K).factorial)) ^ 2) : IsPrincipalIdealRing (𝓞 K) := by have : 0 < finrank ℚ K := finrank_pos -- Lean needs to know that for positivity to succeed diff --git a/Mathlib/NumberTheory/NumberField/Discriminant/Basic.lean b/Mathlib/NumberTheory/NumberField/Discriminant/Basic.lean index 26397e665bada..69a827b17b1c0 100644 --- a/Mathlib/NumberTheory/NumberField/Discriminant/Basic.lean +++ b/Mathlib/NumberTheory/NumberField/Discriminant/Basic.lean @@ -40,7 +40,7 @@ open MeasureTheory MeasureTheory.Measure ZSpan NumberField.mixedEmbedding theorem _root_.NumberField.mixedEmbedding.volume_fundamentalDomain_latticeBasis : volume (fundamentalDomain (latticeBasis K)) = - (2 : ℝ≥0∞)⁻¹ ^ NrComplexPlaces K * sqrt ‖discr K‖₊ := by + (2 : ℝ≥0∞)⁻¹ ^ nrComplexPlaces K * sqrt ‖discr K‖₊ := by let f : Module.Free.ChooseBasisIndex ℤ (𝓞 K) ≃ (K →+* ℂ) := (canonicalEmbedding.latticeBasis K).indexEquiv (Pi.basisFun ℂ _) let e : (index K) ≃ Module.Free.ChooseBasisIndex ℤ (𝓞 K) := (indexEquiv K).trans f.symm @@ -74,7 +74,7 @@ theorem _root_.NumberField.mixedEmbedding.volume_fundamentalDomain_latticeBasis theorem exists_ne_zero_mem_ideal_of_norm_le_mul_sqrt_discr (I : (FractionalIdeal (𝓞 K)⁰ K)ˣ) : ∃ a ∈ (I : FractionalIdeal (𝓞 K)⁰ K), a ≠ 0 ∧ - |Algebra.norm ℚ (a : K)| ≤ FractionalIdeal.absNorm I.1 * (4 / π) ^ NrComplexPlaces K * + |Algebra.norm ℚ (a : K)| ≤ FractionalIdeal.absNorm I.1 * (4 / π) ^ nrComplexPlaces K * (finrank ℚ K).factorial / (finrank ℚ K) ^ (finrank ℚ K) * Real.sqrt |discr K| := by -- The smallest possible value for `exists_ne_zero_mem_ideal_of_norm_le` let B := (minkowskiBound K I * (convexBodySumFactor K)⁻¹).toReal ^ (1 / (finrank ℚ K : ℝ)) @@ -92,8 +92,8 @@ theorem exists_ne_zero_mem_ideal_of_norm_le_mul_sqrt_discr (I : (FractionalIdeal congr 1 rw [eq_comm] calc - _ = FractionalIdeal.absNorm I.1 * (2 : ℝ)⁻¹ ^ NrComplexPlaces K * sqrt ‖discr K‖₊ * - (2 : ℝ) ^ finrank ℚ K * ((2 : ℝ) ^ NrRealPlaces K * (π / 2) ^ NrComplexPlaces K / + _ = FractionalIdeal.absNorm I.1 * (2 : ℝ)⁻¹ ^ nrComplexPlaces K * sqrt ‖discr K‖₊ * + (2 : ℝ) ^ finrank ℚ K * ((2 : ℝ) ^ nrRealPlaces K * (π / 2) ^ nrComplexPlaces K / (Nat.factorial (finrank ℚ K)))⁻¹ := by simp_rw [minkowskiBound, convexBodySumFactor, volume_fundamentalDomain_fractionalIdealLatticeBasis, @@ -102,19 +102,19 @@ theorem exists_ne_zero_mem_ideal_of_norm_le_mul_sqrt_discr (I : (FractionalIdeal rw [ENNReal.toReal_ofReal (Rat.cast_nonneg.mpr (FractionalIdeal.absNorm_nonneg I.1))] simp_rw [NNReal.coe_inv, NNReal.coe_div, NNReal.coe_mul, NNReal.coe_pow, NNReal.coe_div, coe_real_pi, NNReal.coe_ofNat, NNReal.coe_natCast] - _ = FractionalIdeal.absNorm I.1 * (2 : ℝ) ^ (finrank ℚ K - NrComplexPlaces K - NrRealPlaces K + - NrComplexPlaces K : ℤ) * Real.sqrt ‖discr K‖ * Nat.factorial (finrank ℚ K) * - π⁻¹ ^ (NrComplexPlaces K) := by + _ = FractionalIdeal.absNorm I.1 * (2 : ℝ) ^ (finrank ℚ K - nrComplexPlaces K - nrRealPlaces K + + nrComplexPlaces K : ℤ) * Real.sqrt ‖discr K‖ * Nat.factorial (finrank ℚ K) * + π⁻¹ ^ (nrComplexPlaces K) := by simp_rw [inv_div, div_eq_mul_inv, mul_inv, ← zpow_neg_one, ← zpow_natCast, mul_zpow, ← zpow_mul, neg_one_mul, mul_neg_one, neg_neg, Real.coe_sqrt, coe_nnnorm, sub_eq_add_neg, zpow_add₀ (two_ne_zero : (2 : ℝ) ≠ 0)] ring - _ = FractionalIdeal.absNorm I.1 * (2 : ℝ) ^ (2 * NrComplexPlaces K : ℤ) * Real.sqrt ‖discr K‖ * - Nat.factorial (finrank ℚ K) * π⁻¹ ^ (NrComplexPlaces K) := by + _ = FractionalIdeal.absNorm I.1 * (2 : ℝ) ^ (2 * nrComplexPlaces K : ℤ) * Real.sqrt ‖discr K‖ * + Nat.factorial (finrank ℚ K) * π⁻¹ ^ (nrComplexPlaces K) := by congr rw [← card_add_two_mul_card_eq_rank, Nat.cast_add, Nat.cast_mul, Nat.cast_ofNat] ring - _ = FractionalIdeal.absNorm I.1 * (4 / π) ^ NrComplexPlaces K * (finrank ℚ K).factorial * + _ = FractionalIdeal.absNorm I.1 * (4 / π) ^ nrComplexPlaces K * (finrank ℚ K).factorial * Real.sqrt |discr K| := by rw [Int.norm_eq_abs, zpow_mul, show (2 : ℝ) ^ (2 : ℤ) = 4 by norm_cast, div_pow, inv_eq_one_div, div_pow, one_pow, zpow_natCast] @@ -122,7 +122,7 @@ theorem exists_ne_zero_mem_ideal_of_norm_le_mul_sqrt_discr (I : (FractionalIdeal theorem exists_ne_zero_mem_ringOfIntegers_of_norm_le_mul_sqrt_discr : ∃ (a : 𝓞 K), a ≠ 0 ∧ - |Algebra.norm ℚ (a : K)| ≤ (4 / π) ^ NrComplexPlaces K * + |Algebra.norm ℚ (a : K)| ≤ (4 / π) ^ nrComplexPlaces K * (finrank ℚ K).factorial / (finrank ℚ K) ^ (finrank ℚ K) * Real.sqrt |discr K| := by obtain ⟨_, h_mem, h_nz, h_nm⟩ := exists_ne_zero_mem_ideal_of_norm_le_mul_sqrt_discr K ↑1 obtain ⟨a, rfl⟩ := (FractionalIdeal.mem_one_iff _).mp h_mem diff --git a/Mathlib/NumberTheory/NumberField/Embeddings.lean b/Mathlib/NumberTheory/NumberField/Embeddings.lean index d5e8af77ec71d..f041f1b899502 100644 --- a/Mathlib/NumberTheory/NumberField/Embeddings.lean +++ b/Mathlib/NumberTheory/NumberField/Embeddings.lean @@ -551,22 +551,26 @@ section NumberField variable [NumberField K] /-- The number of infinite real places of the number field `K`. -/ -noncomputable abbrev NrRealPlaces := card { w : InfinitePlace K // IsReal w } +noncomputable abbrev nrRealPlaces := card { w : InfinitePlace K // IsReal w } + +@[deprecated (since := "2024-10-24")] alias NrRealPlaces := nrRealPlaces /-- The number of infinite complex places of the number field `K`. -/ -noncomputable abbrev NrComplexPlaces := card { w : InfinitePlace K // IsComplex w } +noncomputable abbrev nrComplexPlaces := card { w : InfinitePlace K // IsComplex w } + +@[deprecated (since := "2024-10-24")] alias NrComplexPlaces := nrComplexPlaces theorem card_real_embeddings : - card { φ : K →+* ℂ // ComplexEmbedding.IsReal φ } = NrRealPlaces K := Fintype.card_congr mkReal + card { φ : K →+* ℂ // ComplexEmbedding.IsReal φ } = nrRealPlaces K := Fintype.card_congr mkReal theorem card_eq_nrRealPlaces_add_nrComplexPlaces : - Fintype.card (InfinitePlace K) = NrRealPlaces K + NrComplexPlaces K := by + Fintype.card (InfinitePlace K) = nrRealPlaces K + nrComplexPlaces K := by convert Fintype.card_subtype_or_disjoint (IsReal (K := K)) (IsComplex (K := K)) (disjoint_isReal_isComplex K) using 1 exact (Fintype.card_of_subtype _ (fun w ↦ ⟨fun _ ↦ isReal_or_isComplex w, fun _ ↦ by simp⟩)).symm theorem card_complex_embeddings : - card { φ : K →+* ℂ // ¬ComplexEmbedding.IsReal φ } = 2 * NrComplexPlaces K := by + card { φ : K →+* ℂ // ¬ComplexEmbedding.IsReal φ } = 2 * nrComplexPlaces K := by suffices ∀ w : { w : InfinitePlace K // IsComplex w }, #{φ : {φ //¬ ComplexEmbedding.IsReal φ} | mkComplex φ = w} = 2 by rw [Fintype.card, Finset.card_eq_sum_ones, ← Finset.sum_fiberwise _ (fun φ => mkComplex φ)] @@ -583,7 +587,7 @@ theorem card_complex_embeddings : · simp_rw [mult, not_isReal_iff_isComplex.mpr hw, ite_false] theorem card_add_two_mul_card_eq_rank : - NrRealPlaces K + 2 * NrComplexPlaces K = finrank ℚ K := by + nrRealPlaces K + 2 * nrComplexPlaces K = finrank ℚ K := by rw [← card_real_embeddings, ← card_complex_embeddings, Fintype.card_subtype_compl, ← Embeddings.card K ℂ, Nat.add_sub_of_le] exact Fintype.card_subtype_le _ @@ -591,10 +595,10 @@ theorem card_add_two_mul_card_eq_rank : variable {K} theorem nrComplexPlaces_eq_zero_of_finrank_eq_one (h : finrank ℚ K = 1) : - NrComplexPlaces K = 0 := by linarith [card_add_two_mul_card_eq_rank K] + nrComplexPlaces K = 0 := by linarith [card_add_two_mul_card_eq_rank K] theorem nrRealPlaces_eq_one_of_finrank_eq_one (h : finrank ℚ K = 1) : - NrRealPlaces K = 1 := by + nrRealPlaces K = 1 := by have := card_add_two_mul_card_eq_rank K rwa [nrComplexPlaces_eq_zero_of_finrank_eq_one h, h, mul_zero, add_zero] at this @@ -1042,7 +1046,7 @@ namespace IsPrimitiveRoot variable {K : Type*} [Field K] [NumberField K] {ζ : K} {k : ℕ} theorem nrRealPlaces_eq_zero_of_two_lt (hk : 2 < k) (hζ : IsPrimitiveRoot ζ k) : - NumberField.InfinitePlace.NrRealPlaces K = 0 := by + NumberField.InfinitePlace.nrRealPlaces K = 0 := by refine (@Fintype.card_eq_zero_iff _ (_)).2 ⟨fun ⟨w, hwreal⟩ ↦ ?_⟩ rw [NumberField.InfinitePlace.isReal_iff] at hwreal let f := w.embedding From b5a5319a32d9785c6c18f40b803ff6af0ff31077 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Fri, 25 Oct 2024 12:58:33 +0000 Subject: [PATCH 415/505] chore(Algebra/Associated): split off Prime and Irreducible (#18224) This PR moves definitions about prime and irreducible elements to their own file. I always found it somewhat weird that associated elements and prime elements had to live together. After this PR: * `Prime`, `Irreducible` and `irreducible_iff_prime` will live in `Algebra/Prime/Defs.lean` * All results on prime and irreducible elements that do not further need associated elements will live in `Algebra/Prime/Lemmas.lean` I considered also splitting up `Associated/Basic.lean` into `Defs`, `Lemmas` and `Prime`, but that might be for a follow-up PR. This is in preparation for further cleaning up `Data/Nat/Prime/Defs.lean`. --- Mathlib.lean | 2 + Mathlib/Algebra/Associated/Basic.lean | 397 +------------------------- Mathlib/Algebra/Prime/Defs.lean | 191 +++++++++++++ Mathlib/Algebra/Prime/Lemmas.lean | 276 ++++++++++++++++++ Mathlib/Data/Nat/Prime/Basic.lean | 5 +- Mathlib/Data/Nat/Prime/Defs.lean | 2 +- Mathlib/RingTheory/Prime.lean | 3 +- 7 files changed, 478 insertions(+), 398 deletions(-) create mode 100644 Mathlib/Algebra/Prime/Defs.lean create mode 100644 Mathlib/Algebra/Prime/Lemmas.lean diff --git a/Mathlib.lean b/Mathlib.lean index a58a3a3939b0d..d885b3af7da93 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -771,6 +771,8 @@ import Mathlib.Algebra.Polynomial.Splits import Mathlib.Algebra.Polynomial.SumIteratedDerivative import Mathlib.Algebra.Polynomial.Taylor import Mathlib.Algebra.Polynomial.UnitTrinomial +import Mathlib.Algebra.Prime.Defs +import Mathlib.Algebra.Prime.Lemmas import Mathlib.Algebra.QuadraticDiscriminant import Mathlib.Algebra.Quandle import Mathlib.Algebra.Quaternion diff --git a/Mathlib/Algebra/Associated/Basic.lean b/Mathlib/Algebra/Associated/Basic.lean index ae0f10085a509..3697b4cbe06ce 100644 --- a/Mathlib/Algebra/Associated/Basic.lean +++ b/Mathlib/Algebra/Associated/Basic.lean @@ -10,19 +10,12 @@ import Mathlib.Algebra.Group.Commute.Units import Mathlib.Algebra.Group.Units.Equiv import Mathlib.Order.BoundedOrder import Mathlib.Algebra.Ring.Units +import Mathlib.Algebra.Prime.Lemmas /-! -# Associated, prime, and irreducible elements. +# Associated elements. -In this file we define the predicate `Prime p` -saying that an element of a commutative monoid with zero is prime. -Namely, `Prime p` means that `p` isn't zero, it isn't a unit, -and `p ∣ a * b → p ∣ a ∨ p ∣ b` for all `a`, `b`; - -In decomposition monoids (e.g., `ℕ`, `ℤ`), this predicate is equivalent to `Irreducible`, -however this is not true in general. - -We also define an equivalence relation `Associated` +In this file we define an equivalence relation `Associated` saying that two elements of a monoid differ by a multiplication by a unit. Then we show that the quotient type `Associates` is a monoid and prove basic properties of this quotient. @@ -33,357 +26,6 @@ assert_not_exists Multiset variable {M N : Type*} -section Prime - -variable [CommMonoidWithZero M] - -/-- An element `p` of a commutative monoid with zero (e.g., a ring) is called *prime*, -if it's not zero, not a unit, and `p ∣ a * b → p ∣ a ∨ p ∣ b` for all `a`, `b`. -/ -def Prime (p : M) : Prop := - p ≠ 0 ∧ ¬IsUnit p ∧ ∀ a b, p ∣ a * b → p ∣ a ∨ p ∣ b - -namespace Prime - -variable {p : M} (hp : Prime p) -include hp - -theorem ne_zero : p ≠ 0 := - hp.1 - -theorem not_unit : ¬IsUnit p := - hp.2.1 - -theorem not_dvd_one : ¬p ∣ 1 := - mt (isUnit_of_dvd_one ·) hp.not_unit - -theorem ne_one : p ≠ 1 := fun h => hp.2.1 (h.symm ▸ isUnit_one) - -theorem dvd_or_dvd {a b : M} (h : p ∣ a * b) : p ∣ a ∨ p ∣ b := - hp.2.2 a b h - -theorem dvd_mul {a b : M} : p ∣ a * b ↔ p ∣ a ∨ p ∣ b := - ⟨hp.dvd_or_dvd, (Or.elim · (dvd_mul_of_dvd_left · _) (dvd_mul_of_dvd_right · _))⟩ - -theorem isPrimal : IsPrimal p := fun _a _b dvd ↦ (hp.dvd_or_dvd dvd).elim - (fun h ↦ ⟨p, 1, h, one_dvd _, (mul_one p).symm⟩) fun h ↦ ⟨1, p, one_dvd _, h, (one_mul p).symm⟩ - -theorem not_dvd_mul {a b : M} (ha : ¬ p ∣ a) (hb : ¬ p ∣ b) : ¬ p ∣ a * b := - hp.dvd_mul.not.mpr <| not_or.mpr ⟨ha, hb⟩ - -theorem dvd_of_dvd_pow {a : M} {n : ℕ} (h : p ∣ a ^ n) : p ∣ a := by - induction n with - | zero => - rw [pow_zero] at h - have := isUnit_of_dvd_one h - have := not_unit hp - contradiction - | succ n ih => - rw [pow_succ'] at h - rcases dvd_or_dvd hp h with dvd_a | dvd_pow - · assumption - · exact ih dvd_pow - -theorem dvd_pow_iff_dvd {a : M} {n : ℕ} (hn : n ≠ 0) : p ∣ a ^ n ↔ p ∣ a := - ⟨hp.dvd_of_dvd_pow, (dvd_pow · hn)⟩ - -end Prime - -@[simp] -theorem not_prime_zero : ¬Prime (0 : M) := fun h => h.ne_zero rfl - -@[simp] -theorem not_prime_one : ¬Prime (1 : M) := fun h => h.not_unit isUnit_one - -section Map - -variable [CommMonoidWithZero N] {F : Type*} {G : Type*} [FunLike F M N] -variable [MonoidWithZeroHomClass F M N] [FunLike G N M] [MulHomClass G N M] -variable (f : F) (g : G) {p : M} - -theorem comap_prime (hinv : ∀ a, g (f a : N) = a) (hp : Prime (f p)) : Prime p := - ⟨fun h => hp.1 <| by simp [h], fun h => hp.2.1 <| h.map f, fun a b h => by - refine - (hp.2.2 (f a) (f b) <| by - convert map_dvd f h - simp).imp - ?_ ?_ <;> - · intro h - convert ← map_dvd g h <;> apply hinv⟩ - -theorem MulEquiv.prime_iff (e : M ≃* N) : Prime p ↔ Prime (e p) := - ⟨fun h => (comap_prime e.symm e fun a => by simp) <| (e.symm_apply_apply p).substr h, - comap_prime e e.symm fun a => by simp⟩ - -end Map - -end Prime - -theorem Prime.left_dvd_or_dvd_right_of_dvd_mul [CancelCommMonoidWithZero M] {p : M} (hp : Prime p) - {a b : M} : a ∣ p * b → p ∣ a ∨ a ∣ b := by - rintro ⟨c, hc⟩ - rcases hp.2.2 a c (hc ▸ dvd_mul_right _ _) with (h | ⟨x, rfl⟩) - · exact Or.inl h - · rw [mul_left_comm, mul_right_inj' hp.ne_zero] at hc - exact Or.inr (hc.symm ▸ dvd_mul_right _ _) - -theorem Prime.pow_dvd_of_dvd_mul_left [CancelCommMonoidWithZero M] {p a b : M} (hp : Prime p) - (n : ℕ) (h : ¬p ∣ a) (h' : p ^ n ∣ a * b) : p ^ n ∣ b := by - induction n with - | zero => - rw [pow_zero] - exact one_dvd b - | succ n ih => - obtain ⟨c, rfl⟩ := ih (dvd_trans (pow_dvd_pow p n.le_succ) h') - rw [pow_succ] - apply mul_dvd_mul_left _ ((hp.dvd_or_dvd _).resolve_left h) - rwa [← mul_dvd_mul_iff_left (pow_ne_zero n hp.ne_zero), ← pow_succ, mul_left_comm] - -theorem Prime.pow_dvd_of_dvd_mul_right [CancelCommMonoidWithZero M] {p a b : M} (hp : Prime p) - (n : ℕ) (h : ¬p ∣ b) (h' : p ^ n ∣ a * b) : p ^ n ∣ a := by - rw [mul_comm] at h' - exact hp.pow_dvd_of_dvd_mul_left n h h' - -theorem Prime.dvd_of_pow_dvd_pow_mul_pow_of_square_not_dvd [CancelCommMonoidWithZero M] {p a b : M} - {n : ℕ} (hp : Prime p) (hpow : p ^ n.succ ∣ a ^ n.succ * b ^ n) (hb : ¬p ^ 2 ∣ b) : p ∣ a := by - -- Suppose `p ∣ b`, write `b = p * x` and `hy : a ^ n.succ * b ^ n = p ^ n.succ * y`. - rcases hp.dvd_or_dvd ((dvd_pow_self p (Nat.succ_ne_zero n)).trans hpow) with H | hbdiv - · exact hp.dvd_of_dvd_pow H - obtain ⟨x, rfl⟩ := hp.dvd_of_dvd_pow hbdiv - obtain ⟨y, hy⟩ := hpow - -- Then we can divide out a common factor of `p ^ n` from the equation `hy`. - have : a ^ n.succ * x ^ n = p * y := by - refine mul_left_cancel₀ (pow_ne_zero n hp.ne_zero) ?_ - rw [← mul_assoc _ p, ← pow_succ, ← hy, mul_pow, ← mul_assoc (a ^ n.succ), mul_comm _ (p ^ n), - mul_assoc] - -- So `p ∣ a` (and we're done) or `p ∣ x`, which can't be the case since it implies `p^2 ∣ b`. - refine hp.dvd_of_dvd_pow ((hp.dvd_or_dvd ⟨_, this⟩).resolve_right fun hdvdx => hb ?_) - obtain ⟨z, rfl⟩ := hp.dvd_of_dvd_pow hdvdx - rw [pow_two, ← mul_assoc] - exact dvd_mul_right _ _ - -theorem prime_pow_succ_dvd_mul {M : Type*} [CancelCommMonoidWithZero M] {p x y : M} (h : Prime p) - {i : ℕ} (hxy : p ^ (i + 1) ∣ x * y) : p ^ (i + 1) ∣ x ∨ p ∣ y := by - rw [or_iff_not_imp_right] - intro hy - induction i generalizing x with - | zero => rw [pow_one] at hxy ⊢; exact (h.dvd_or_dvd hxy).resolve_right hy - | succ i ih => - rw [pow_succ'] at hxy ⊢ - obtain ⟨x', rfl⟩ := (h.dvd_or_dvd (dvd_of_mul_right_dvd hxy)).resolve_right hy - rw [mul_assoc] at hxy - exact mul_dvd_mul_left p (ih ((mul_dvd_mul_iff_left h.ne_zero).mp hxy)) - -/-- `Irreducible p` states that `p` is non-unit and only factors into units. - -We explicitly avoid stating that `p` is non-zero, this would require a semiring. Assuming only a -monoid allows us to reuse irreducible for associated elements. --/ -structure Irreducible [Monoid M] (p : M) : Prop where - /-- `p` is not a unit -/ - not_unit : ¬IsUnit p - /-- if `p` factors then one factor is a unit -/ - isUnit_or_isUnit' : ∀ a b, p = a * b → IsUnit a ∨ IsUnit b - -namespace Irreducible - -theorem not_dvd_one [CommMonoid M] {p : M} (hp : Irreducible p) : ¬p ∣ 1 := - mt (isUnit_of_dvd_one ·) hp.not_unit - -theorem isUnit_or_isUnit [Monoid M] {p : M} (hp : Irreducible p) {a b : M} (h : p = a * b) : - IsUnit a ∨ IsUnit b := - hp.isUnit_or_isUnit' a b h - -end Irreducible - -theorem irreducible_iff [Monoid M] {p : M} : - Irreducible p ↔ ¬IsUnit p ∧ ∀ a b, p = a * b → IsUnit a ∨ IsUnit b := - ⟨fun h => ⟨h.1, h.2⟩, fun h => ⟨h.1, h.2⟩⟩ - -@[simp] -theorem not_irreducible_one [Monoid M] : ¬Irreducible (1 : M) := by simp [irreducible_iff] - -theorem Irreducible.ne_one [Monoid M] : ∀ {p : M}, Irreducible p → p ≠ 1 - | _, hp, rfl => not_irreducible_one hp - -@[simp] -theorem not_irreducible_zero [MonoidWithZero M] : ¬Irreducible (0 : M) - | ⟨hn0, h⟩ => - have : IsUnit (0 : M) ∨ IsUnit (0 : M) := h 0 0 (mul_zero 0).symm - this.elim hn0 hn0 - -theorem Irreducible.ne_zero [MonoidWithZero M] : ∀ {p : M}, Irreducible p → p ≠ 0 - | _, hp, rfl => not_irreducible_zero hp - -theorem of_irreducible_mul {M} [Monoid M] {x y : M} : Irreducible (x * y) → IsUnit x ∨ IsUnit y - | ⟨_, h⟩ => h _ _ rfl - -theorem not_irreducible_pow {M} [Monoid M] {x : M} {n : ℕ} (hn : n ≠ 1) : - ¬ Irreducible (x ^ n) := by - cases n with - | zero => simp - | succ n => - intro ⟨h₁, h₂⟩ - have := h₂ _ _ (pow_succ _ _) - rw [isUnit_pow_iff (Nat.succ_ne_succ.mp hn), or_self] at this - exact h₁ (this.pow _) - -theorem irreducible_or_factor {M} [Monoid M] (x : M) (h : ¬IsUnit x) : - Irreducible x ∨ ∃ a b, ¬IsUnit a ∧ ¬IsUnit b ∧ a * b = x := by - haveI := Classical.dec - refine or_iff_not_imp_right.2 fun H => ?_ - simp? [h, irreducible_iff] at H ⊢ says - simp only [exists_and_left, not_exists, not_and, irreducible_iff, h, not_false_eq_true, - true_and] at H ⊢ - refine fun a b h => by_contradiction fun o => ?_ - simp? [not_or] at o says simp only [not_or] at o - exact H _ o.1 _ o.2 h.symm - -/-- If `p` and `q` are irreducible, then `p ∣ q` implies `q ∣ p`. -/ -theorem Irreducible.dvd_symm [Monoid M] {p q : M} (hp : Irreducible p) (hq : Irreducible q) : - p ∣ q → q ∣ p := by - rintro ⟨q', rfl⟩ - rw [IsUnit.mul_right_dvd (Or.resolve_left (of_irreducible_mul hq) hp.not_unit)] - -theorem Irreducible.dvd_comm [Monoid M] {p q : M} (hp : Irreducible p) (hq : Irreducible q) : - p ∣ q ↔ q ∣ p := - ⟨hp.dvd_symm hq, hq.dvd_symm hp⟩ - -theorem Irreducible.of_map {F : Type*} [Monoid M] [Monoid N] [FunLike F M N] [MonoidHomClass F M N] - {f : F} [IsLocalHom f] {x} (hfx : Irreducible (f x)) : Irreducible x := - ⟨fun hu ↦ hfx.not_unit <| hu.map f, - by rintro p q rfl - exact (hfx.isUnit_or_isUnit <| map_mul f p q).imp (.of_map f _) (.of_map f _)⟩ - -section - -variable [Monoid M] - -theorem irreducible_units_mul (a : Mˣ) (b : M) : Irreducible (↑a * b) ↔ Irreducible b := by - simp only [irreducible_iff, Units.isUnit_units_mul, and_congr_right_iff] - refine fun _ => ⟨fun h A B HAB => ?_, fun h A B HAB => ?_⟩ - · rw [← a.isUnit_units_mul] - apply h - rw [mul_assoc, ← HAB] - · rw [← a⁻¹.isUnit_units_mul] - apply h - rw [mul_assoc, ← HAB, Units.inv_mul_cancel_left] - -theorem irreducible_isUnit_mul {a b : M} (h : IsUnit a) : Irreducible (a * b) ↔ Irreducible b := - let ⟨a, ha⟩ := h - ha ▸ irreducible_units_mul a b - -theorem irreducible_mul_units (a : Mˣ) (b : M) : Irreducible (b * ↑a) ↔ Irreducible b := by - simp only [irreducible_iff, Units.isUnit_mul_units, and_congr_right_iff] - refine fun _ => ⟨fun h A B HAB => ?_, fun h A B HAB => ?_⟩ - · rw [← Units.isUnit_mul_units B a] - apply h - rw [← mul_assoc, ← HAB] - · rw [← Units.isUnit_mul_units B a⁻¹] - apply h - rw [← mul_assoc, ← HAB, Units.mul_inv_cancel_right] - -theorem irreducible_mul_isUnit {a b : M} (h : IsUnit a) : Irreducible (b * a) ↔ Irreducible b := - let ⟨a, ha⟩ := h - ha ▸ irreducible_mul_units a b - -theorem irreducible_mul_iff {a b : M} : - Irreducible (a * b) ↔ Irreducible a ∧ IsUnit b ∨ Irreducible b ∧ IsUnit a := by - constructor - · refine fun h => Or.imp (fun h' => ⟨?_, h'⟩) (fun h' => ⟨?_, h'⟩) (h.isUnit_or_isUnit rfl).symm - · rwa [irreducible_mul_isUnit h'] at h - · rwa [irreducible_isUnit_mul h'] at h - · rintro (⟨ha, hb⟩ | ⟨hb, ha⟩) - · rwa [irreducible_mul_isUnit hb] - · rwa [irreducible_isUnit_mul ha] - -variable [Monoid N] {F : Type*} [EquivLike F M N] [MulEquivClass F M N] (f : F) - -open MulEquiv - -/-- -Irreducibility is preserved by multiplicative equivalences. -Note that surjective + local hom is not enough. Consider the additive monoids `M = ℕ ⊕ ℕ`, `N = ℕ`, -with a surjective local (additive) hom `f : M →+ N` sending `(m, n)` to `2m + n`. -It is local because the only add unit in `N` is `0`, with preimage `{(0, 0)}` also an add unit. -Then `x = (1, 0)` is irreducible in `M`, but `f x = 2 = 1 + 1` is not irreducible in `N`. --/ -theorem Irreducible.map {x : M} (h : Irreducible x) : Irreducible (f x) := - ⟨fun g ↦ h.not_unit g.of_map, fun a b g ↦ - let f := MulEquivClass.toMulEquiv f - (h.isUnit_or_isUnit (symm_apply_apply f x ▸ map_mul f.symm a b ▸ congrArg f.symm g)).imp - (·.of_map) (·.of_map)⟩ - -theorem MulEquiv.irreducible_iff (f : F) {a : M} : - Irreducible (f a) ↔ Irreducible a := - ⟨Irreducible.of_map, Irreducible.map f⟩ - -end - -section CommMonoid - -variable [CommMonoid M] {a : M} - -theorem Irreducible.not_square (ha : Irreducible a) : ¬IsSquare a := by - rw [isSquare_iff_exists_sq] - rintro ⟨b, rfl⟩ - exact not_irreducible_pow (by decide) ha - -theorem IsSquare.not_irreducible (ha : IsSquare a) : ¬Irreducible a := fun h => h.not_square ha - -end CommMonoid - -section CommMonoidWithZero - -variable [CommMonoidWithZero M] - -theorem Irreducible.prime_of_isPrimal {a : M} - (irr : Irreducible a) (primal : IsPrimal a) : Prime a := - ⟨irr.ne_zero, irr.not_unit, fun a b dvd ↦ by - obtain ⟨d₁, d₂, h₁, h₂, rfl⟩ := primal dvd - exact (of_irreducible_mul irr).symm.imp (·.mul_right_dvd.mpr h₁) (·.mul_left_dvd.mpr h₂)⟩ - -theorem Irreducible.prime [DecompositionMonoid M] {a : M} (irr : Irreducible a) : Prime a := - irr.prime_of_isPrimal (DecompositionMonoid.primal a) - -end CommMonoidWithZero - -section CancelCommMonoidWithZero - -variable [CancelCommMonoidWithZero M] {a p : M} - -protected theorem Prime.irreducible (hp : Prime p) : Irreducible p := - ⟨hp.not_unit, fun a b ↦ by - rintro rfl - exact (hp.dvd_or_dvd dvd_rfl).symm.imp - (isUnit_of_dvd_one <| (mul_dvd_mul_iff_right <| right_ne_zero_of_mul hp.ne_zero).mp <| - dvd_mul_of_dvd_right · _) - (isUnit_of_dvd_one <| (mul_dvd_mul_iff_left <| left_ne_zero_of_mul hp.ne_zero).mp <| - dvd_mul_of_dvd_left · _)⟩ - -theorem irreducible_iff_prime [DecompositionMonoid M] {a : M} : Irreducible a ↔ Prime a := - ⟨Irreducible.prime, Prime.irreducible⟩ - -theorem succ_dvd_or_succ_dvd_of_succ_sum_dvd_mul (hp : Prime p) {a b : M} {k l : ℕ} : - p ^ k ∣ a → p ^ l ∣ b → p ^ (k + l + 1) ∣ a * b → p ^ (k + 1) ∣ a ∨ p ^ (l + 1) ∣ b := - fun ⟨x, hx⟩ ⟨y, hy⟩ ⟨z, hz⟩ => - have h : p ^ (k + l) * (x * y) = p ^ (k + l) * (p * z) := by - simpa [mul_comm, pow_add, hx, hy, mul_assoc, mul_left_comm] using hz - have hp0 : p ^ (k + l) ≠ 0 := pow_ne_zero _ hp.ne_zero - have hpd : p ∣ x * y := ⟨z, by rwa [mul_right_inj' hp0] at h⟩ - (hp.dvd_or_dvd hpd).elim - (fun ⟨d, hd⟩ => Or.inl ⟨d, by simp [*, pow_succ, mul_comm, mul_left_comm, mul_assoc]⟩) - fun ⟨d, hd⟩ => Or.inr ⟨d, by simp [*, pow_succ, mul_comm, mul_left_comm, mul_assoc]⟩ - -theorem Prime.not_square (hp : Prime p) : ¬IsSquare p := - hp.irreducible.not_square - -theorem IsSquare.not_prime (ha : IsSquare a) : ¬Prime a := fun h => h.not_square ha - -theorem not_prime_pow {n : ℕ} (hn : n ≠ 1) : ¬Prime (a ^ n) := fun hp => - not_irreducible_pow hn hp.irreducible - -end CancelCommMonoidWithZero - /-- Two elements of a `Monoid` are `Associated` if one of them is another one multiplied by a unit on the right. -/ def Associated [Monoid M] (x y : M) : Prop := @@ -1107,19 +749,6 @@ end Associates section CommMonoidWithZero -theorem DvdNotUnit.isUnit_of_irreducible_right [CommMonoidWithZero M] {p q : M} - (h : DvdNotUnit p q) (hq : Irreducible q) : IsUnit p := by - obtain ⟨_, x, hx, hx'⟩ := h - exact Or.resolve_right ((irreducible_iff.1 hq).right p x hx') hx - -theorem not_irreducible_of_not_unit_dvdNotUnit [CommMonoidWithZero M] {p q : M} (hp : ¬IsUnit p) - (h : DvdNotUnit p q) : ¬Irreducible q := - mt h.isUnit_of_irreducible_right hp - -theorem DvdNotUnit.not_unit [CommMonoidWithZero M] {p q : M} (hp : DvdNotUnit p q) : ¬IsUnit q := by - obtain ⟨-, x, hx, rfl⟩ := hp - exact fun hc => hx (isUnit_iff_dvd_one.mpr (dvd_of_mul_left_dvd (isUnit_iff_dvd_one.mp hc))) - theorem dvdNotUnit_of_dvdNotUnit_associated [CommMonoidWithZero M] [Nontrivial M] {p q r : M} (h : DvdNotUnit p q) (h' : Associated q r) : DvdNotUnit p r := by obtain ⟨u, rfl⟩ := Associated.symm h' @@ -1144,26 +773,6 @@ theorem DvdNotUnit.not_associated [CancelCommMonoidWithZero M] {p q : M} (h : Dv rcases (mul_right_inj' hp).mp hx' with rfl exact hx a.isUnit -theorem DvdNotUnit.ne [CancelCommMonoidWithZero M] {p q : M} (h : DvdNotUnit p q) : p ≠ q := by - by_contra hcontra - obtain ⟨hp, x, hx', hx''⟩ := h - conv_lhs at hx'' => rw [← hcontra, ← mul_one p] - rw [(mul_left_cancel₀ hp hx'').symm] at hx' - exact hx' isUnit_one - -theorem pow_injective_of_not_isUnit [CancelCommMonoidWithZero M] {q : M} (hq : ¬IsUnit q) - (hq' : q ≠ 0) : Function.Injective fun n : ℕ => q ^ n := by - refine injective_of_lt_imp_ne fun n m h => DvdNotUnit.ne ⟨pow_ne_zero n hq', q ^ (m - n), ?_, ?_⟩ - · exact not_isUnit_of_not_isUnit_dvd hq (dvd_pow (dvd_refl _) (Nat.sub_pos_of_lt h).ne') - · exact (pow_mul_pow_sub q h.le).symm - -@[deprecated (since := "2024-09-22")] -alias pow_injective_of_not_unit := pow_injective_of_not_isUnit - -theorem pow_inj_of_not_isUnit [CancelCommMonoidWithZero M] {q : M} (hq : ¬IsUnit q) - (hq' : q ≠ 0) {m n : ℕ} : q ^ m = q ^ n ↔ m = n := - (pow_injective_of_not_isUnit hq hq').eq_iff - theorem dvd_prime_pow [CancelCommMonoidWithZero M] {p q : M} (hp : Prime p) (n : ℕ) : q ∣ p ^ n ↔ ∃ i ≤ n, Associated q (p ^ i) := by induction n generalizing q with diff --git a/Mathlib/Algebra/Prime/Defs.lean b/Mathlib/Algebra/Prime/Defs.lean new file mode 100644 index 0000000000000..90165b72d0050 --- /dev/null +++ b/Mathlib/Algebra/Prime/Defs.lean @@ -0,0 +1,191 @@ +/- +Copyright (c) 2018 Johannes Hölzl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johannes Hölzl, Jens Wagemaker +-/ +import Mathlib.Algebra.GroupWithZero.Divisibility + +/-! +# Prime and irreducible elements. + +In this file we define the predicate `Prime p` +saying that an element of a commutative monoid with zero is prime. +Namely, `Prime p` means that `p` isn't zero, it isn't a unit, +and `p ∣ a * b → p ∣ a ∨ p ∣ b` for all `a`, `b`; + +In decomposition monoids (e.g., `ℕ`, `ℤ`), this predicate is equivalent to `Irreducible` +(see `irreducible_iff_prime`), however this is not true in general. + +## Main definitions + * `Prime`: a prime element of a commutative monoid with zero + * `Irreducible`: an irreducible element of a commutative monoid with zero + +## Main results + * `irreducible_iff_prime`: the two definitions are equivalent in a decomposition monoid. +-/ + +assert_not_exists OrderedCommMonoid +assert_not_exists Multiset + +variable {M N : Type*} + +section Prime + +variable [CommMonoidWithZero M] + +/-- An element `p` of a commutative monoid with zero (e.g., a ring) is called *prime*, +if it's not zero, not a unit, and `p ∣ a * b → p ∣ a ∨ p ∣ b` for all `a`, `b`. -/ +def Prime (p : M) : Prop := + p ≠ 0 ∧ ¬IsUnit p ∧ ∀ a b, p ∣ a * b → p ∣ a ∨ p ∣ b + +namespace Prime + +variable {p : M} (hp : Prime p) +include hp + +theorem ne_zero : p ≠ 0 := + hp.1 + +theorem not_unit : ¬IsUnit p := + hp.2.1 + +theorem not_dvd_one : ¬p ∣ 1 := + mt (isUnit_of_dvd_one ·) hp.not_unit + +theorem ne_one : p ≠ 1 := fun h => hp.2.1 (h.symm ▸ isUnit_one) + +theorem dvd_or_dvd {a b : M} (h : p ∣ a * b) : p ∣ a ∨ p ∣ b := + hp.2.2 a b h + +theorem dvd_mul {a b : M} : p ∣ a * b ↔ p ∣ a ∨ p ∣ b := + ⟨hp.dvd_or_dvd, (Or.elim · (dvd_mul_of_dvd_left · _) (dvd_mul_of_dvd_right · _))⟩ + +theorem isPrimal : IsPrimal p := fun _a _b dvd ↦ (hp.dvd_or_dvd dvd).elim + (fun h ↦ ⟨p, 1, h, one_dvd _, (mul_one p).symm⟩) fun h ↦ ⟨1, p, one_dvd _, h, (one_mul p).symm⟩ + +theorem not_dvd_mul {a b : M} (ha : ¬ p ∣ a) (hb : ¬ p ∣ b) : ¬ p ∣ a * b := + hp.dvd_mul.not.mpr <| not_or.mpr ⟨ha, hb⟩ + +theorem dvd_of_dvd_pow {a : M} {n : ℕ} (h : p ∣ a ^ n) : p ∣ a := by + induction n with + | zero => + rw [pow_zero] at h + have := isUnit_of_dvd_one h + have := not_unit hp + contradiction + | succ n ih => + rw [pow_succ'] at h + rcases dvd_or_dvd hp h with dvd_a | dvd_pow + · assumption + · exact ih dvd_pow + +theorem dvd_pow_iff_dvd {a : M} {n : ℕ} (hn : n ≠ 0) : p ∣ a ^ n ↔ p ∣ a := + ⟨hp.dvd_of_dvd_pow, (dvd_pow · hn)⟩ + +end Prime + +@[simp] +theorem not_prime_zero : ¬Prime (0 : M) := fun h => h.ne_zero rfl + +@[simp] +theorem not_prime_one : ¬Prime (1 : M) := fun h => h.not_unit isUnit_one + +end Prime + +/-- `Irreducible p` states that `p` is non-unit and only factors into units. + +We explicitly avoid stating that `p` is non-zero, this would require a semiring. Assuming only a +monoid allows us to reuse irreducible for associated elements. +-/ +structure Irreducible [Monoid M] (p : M) : Prop where + /-- `p` is not a unit -/ + not_unit : ¬IsUnit p + /-- if `p` factors then one factor is a unit -/ + isUnit_or_isUnit' : ∀ a b, p = a * b → IsUnit a ∨ IsUnit b + +namespace Irreducible + +theorem not_dvd_one [CommMonoid M] {p : M} (hp : Irreducible p) : ¬p ∣ 1 := + mt (isUnit_of_dvd_one ·) hp.not_unit + +theorem isUnit_or_isUnit [Monoid M] {p : M} (hp : Irreducible p) {a b : M} (h : p = a * b) : + IsUnit a ∨ IsUnit b := + hp.isUnit_or_isUnit' a b h + +end Irreducible + +theorem irreducible_iff [Monoid M] {p : M} : + Irreducible p ↔ ¬IsUnit p ∧ ∀ a b, p = a * b → IsUnit a ∨ IsUnit b := + ⟨fun h => ⟨h.1, h.2⟩, fun h => ⟨h.1, h.2⟩⟩ + +@[simp] +theorem not_irreducible_one [Monoid M] : ¬Irreducible (1 : M) := by simp [irreducible_iff] + +theorem Irreducible.ne_one [Monoid M] : ∀ {p : M}, Irreducible p → p ≠ 1 + | _, hp, rfl => not_irreducible_one hp + +@[simp] +theorem not_irreducible_zero [MonoidWithZero M] : ¬Irreducible (0 : M) + | ⟨hn0, h⟩ => + have : IsUnit (0 : M) ∨ IsUnit (0 : M) := h 0 0 (mul_zero 0).symm + this.elim hn0 hn0 + +theorem Irreducible.ne_zero [MonoidWithZero M] : ∀ {p : M}, Irreducible p → p ≠ 0 + | _, hp, rfl => not_irreducible_zero hp + +theorem of_irreducible_mul {M} [Monoid M] {x y : M} : Irreducible (x * y) → IsUnit x ∨ IsUnit y + | ⟨_, h⟩ => h _ _ rfl + +theorem irreducible_or_factor {M} [Monoid M] (x : M) (h : ¬IsUnit x) : + Irreducible x ∨ ∃ a b, ¬IsUnit a ∧ ¬IsUnit b ∧ a * b = x := by + haveI := Classical.dec + refine or_iff_not_imp_right.2 fun H => ?_ + simp? [h, irreducible_iff] at H ⊢ says + simp only [exists_and_left, not_exists, not_and, irreducible_iff, h, not_false_eq_true, + true_and] at H ⊢ + refine fun a b h => by_contradiction fun o => ?_ + simp? [not_or] at o says simp only [not_or] at o + exact H _ o.1 _ o.2 h.symm + +/-- If `p` and `q` are irreducible, then `p ∣ q` implies `q ∣ p`. -/ +theorem Irreducible.dvd_symm [Monoid M] {p q : M} (hp : Irreducible p) (hq : Irreducible q) : + p ∣ q → q ∣ p := by + rintro ⟨q', rfl⟩ + rw [IsUnit.mul_right_dvd (Or.resolve_left (of_irreducible_mul hq) hp.not_unit)] + +theorem Irreducible.dvd_comm [Monoid M] {p q : M} (hp : Irreducible p) (hq : Irreducible q) : + p ∣ q ↔ q ∣ p := + ⟨hp.dvd_symm hq, hq.dvd_symm hp⟩ + +section CommMonoidWithZero + +variable [CommMonoidWithZero M] + +theorem Irreducible.prime_of_isPrimal {a : M} + (irr : Irreducible a) (primal : IsPrimal a) : Prime a := + ⟨irr.ne_zero, irr.not_unit, fun a b dvd ↦ by + obtain ⟨d₁, d₂, h₁, h₂, rfl⟩ := primal dvd + exact (of_irreducible_mul irr).symm.imp (·.mul_right_dvd.mpr h₁) (·.mul_left_dvd.mpr h₂)⟩ + +theorem Irreducible.prime [DecompositionMonoid M] {a : M} (irr : Irreducible a) : Prime a := + irr.prime_of_isPrimal (DecompositionMonoid.primal a) + +end CommMonoidWithZero + +section CancelCommMonoidWithZero + +variable [CancelCommMonoidWithZero M] {a p : M} + +protected theorem Prime.irreducible (hp : Prime p) : Irreducible p := + ⟨hp.not_unit, fun a b ↦ by + rintro rfl + exact (hp.dvd_or_dvd dvd_rfl).symm.imp + (isUnit_of_dvd_one <| (mul_dvd_mul_iff_right <| right_ne_zero_of_mul hp.ne_zero).mp <| + dvd_mul_of_dvd_right · _) + (isUnit_of_dvd_one <| (mul_dvd_mul_iff_left <| left_ne_zero_of_mul hp.ne_zero).mp <| + dvd_mul_of_dvd_left · _)⟩ + +theorem irreducible_iff_prime [DecompositionMonoid M] {a : M} : Irreducible a ↔ Prime a := + ⟨Irreducible.prime, Prime.irreducible⟩ + +end CancelCommMonoidWithZero diff --git a/Mathlib/Algebra/Prime/Lemmas.lean b/Mathlib/Algebra/Prime/Lemmas.lean new file mode 100644 index 0000000000000..96ee461f3831f --- /dev/null +++ b/Mathlib/Algebra/Prime/Lemmas.lean @@ -0,0 +1,276 @@ +/- +Copyright (c) 2018 Johannes Hölzl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johannes Hölzl, Jens Wagemaker +-/ +import Mathlib.Algebra.Group.Commute.Units +import Mathlib.Algebra.Group.Even +import Mathlib.Algebra.Group.Units.Equiv +import Mathlib.Algebra.GroupWithZero.Hom +import Mathlib.Algebra.Prime.Defs +import Mathlib.Order.Lattice + +/-! +# Associated, prime, and irreducible elements. + +In this file we define the predicate `Prime p` +saying that an element of a commutative monoid with zero is prime. +Namely, `Prime p` means that `p` isn't zero, it isn't a unit, +and `p ∣ a * b → p ∣ a ∨ p ∣ b` for all `a`, `b`; + +In decomposition monoids (e.g., `ℕ`, `ℤ`), this predicate is equivalent to `Irreducible`, +however this is not true in general. + +We also define an equivalence relation `Associated` +saying that two elements of a monoid differ by a multiplication by a unit. +Then we show that the quotient type `Associates` is a monoid +and prove basic properties of this quotient. +-/ + +assert_not_exists OrderedCommMonoid +assert_not_exists Multiset + +variable {M N : Type*} + +section Prime + +variable [CommMonoidWithZero M] + +section Map + +variable [CommMonoidWithZero N] {F : Type*} {G : Type*} [FunLike F M N] +variable [MonoidWithZeroHomClass F M N] [FunLike G N M] [MulHomClass G N M] +variable (f : F) (g : G) {p : M} + +theorem comap_prime (hinv : ∀ a, g (f a : N) = a) (hp : Prime (f p)) : Prime p := + ⟨fun h => hp.1 <| by simp [h], fun h => hp.2.1 <| h.map f, fun a b h => by + refine + (hp.2.2 (f a) (f b) <| by + convert map_dvd f h + simp).imp + ?_ ?_ <;> + · intro h + convert ← map_dvd g h <;> apply hinv⟩ + +theorem MulEquiv.prime_iff (e : M ≃* N) : Prime p ↔ Prime (e p) := + ⟨fun h => (comap_prime e.symm e fun a => by simp) <| (e.symm_apply_apply p).substr h, + comap_prime e e.symm fun a => by simp⟩ + +end Map + +end Prime + +theorem Prime.left_dvd_or_dvd_right_of_dvd_mul [CancelCommMonoidWithZero M] {p : M} (hp : Prime p) + {a b : M} : a ∣ p * b → p ∣ a ∨ a ∣ b := by + rintro ⟨c, hc⟩ + rcases hp.2.2 a c (hc ▸ dvd_mul_right _ _) with (h | ⟨x, rfl⟩) + · exact Or.inl h + · rw [mul_left_comm, mul_right_inj' hp.ne_zero] at hc + exact Or.inr (hc.symm ▸ dvd_mul_right _ _) + +theorem Prime.pow_dvd_of_dvd_mul_left [CancelCommMonoidWithZero M] {p a b : M} (hp : Prime p) + (n : ℕ) (h : ¬p ∣ a) (h' : p ^ n ∣ a * b) : p ^ n ∣ b := by + induction n with + | zero => + rw [pow_zero] + exact one_dvd b + | succ n ih => + obtain ⟨c, rfl⟩ := ih (dvd_trans (pow_dvd_pow p n.le_succ) h') + rw [pow_succ] + apply mul_dvd_mul_left _ ((hp.dvd_or_dvd _).resolve_left h) + rwa [← mul_dvd_mul_iff_left (pow_ne_zero n hp.ne_zero), ← pow_succ, mul_left_comm] + +theorem Prime.pow_dvd_of_dvd_mul_right [CancelCommMonoidWithZero M] {p a b : M} (hp : Prime p) + (n : ℕ) (h : ¬p ∣ b) (h' : p ^ n ∣ a * b) : p ^ n ∣ a := by + rw [mul_comm] at h' + exact hp.pow_dvd_of_dvd_mul_left n h h' + +theorem Prime.dvd_of_pow_dvd_pow_mul_pow_of_square_not_dvd [CancelCommMonoidWithZero M] {p a b : M} + {n : ℕ} (hp : Prime p) (hpow : p ^ n.succ ∣ a ^ n.succ * b ^ n) (hb : ¬p ^ 2 ∣ b) : p ∣ a := by + -- Suppose `p ∣ b`, write `b = p * x` and `hy : a ^ n.succ * b ^ n = p ^ n.succ * y`. + rcases hp.dvd_or_dvd ((dvd_pow_self p (Nat.succ_ne_zero n)).trans hpow) with H | hbdiv + · exact hp.dvd_of_dvd_pow H + obtain ⟨x, rfl⟩ := hp.dvd_of_dvd_pow hbdiv + obtain ⟨y, hy⟩ := hpow + -- Then we can divide out a common factor of `p ^ n` from the equation `hy`. + have : a ^ n.succ * x ^ n = p * y := by + refine mul_left_cancel₀ (pow_ne_zero n hp.ne_zero) ?_ + rw [← mul_assoc _ p, ← pow_succ, ← hy, mul_pow, ← mul_assoc (a ^ n.succ), mul_comm _ (p ^ n), + mul_assoc] + -- So `p ∣ a` (and we're done) or `p ∣ x`, which can't be the case since it implies `p^2 ∣ b`. + refine hp.dvd_of_dvd_pow ((hp.dvd_or_dvd ⟨_, this⟩).resolve_right fun hdvdx => hb ?_) + obtain ⟨z, rfl⟩ := hp.dvd_of_dvd_pow hdvdx + rw [pow_two, ← mul_assoc] + exact dvd_mul_right _ _ + +theorem prime_pow_succ_dvd_mul {M : Type*} [CancelCommMonoidWithZero M] {p x y : M} (h : Prime p) + {i : ℕ} (hxy : p ^ (i + 1) ∣ x * y) : p ^ (i + 1) ∣ x ∨ p ∣ y := by + rw [or_iff_not_imp_right] + intro hy + induction i generalizing x with + | zero => rw [pow_one] at hxy ⊢; exact (h.dvd_or_dvd hxy).resolve_right hy + | succ i ih => + rw [pow_succ'] at hxy ⊢ + obtain ⟨x', rfl⟩ := (h.dvd_or_dvd (dvd_of_mul_right_dvd hxy)).resolve_right hy + rw [mul_assoc] at hxy + exact mul_dvd_mul_left p (ih ((mul_dvd_mul_iff_left h.ne_zero).mp hxy)) + +theorem not_irreducible_pow {M} [Monoid M] {x : M} {n : ℕ} (hn : n ≠ 1) : + ¬ Irreducible (x ^ n) := by + cases n with + | zero => simp + | succ n => + intro ⟨h₁, h₂⟩ + have := h₂ _ _ (pow_succ _ _) + rw [isUnit_pow_iff (Nat.succ_ne_succ.mp hn), or_self] at this + exact h₁ (this.pow _) + +theorem Irreducible.of_map {F : Type*} [Monoid M] [Monoid N] [FunLike F M N] [MonoidHomClass F M N] + {f : F} [IsLocalHom f] {x} (hfx : Irreducible (f x)) : Irreducible x := + ⟨fun hu ↦ hfx.not_unit <| hu.map f, + by rintro p q rfl + exact (hfx.isUnit_or_isUnit <| map_mul f p q).imp (.of_map f _) (.of_map f _)⟩ + +section + +variable [Monoid M] + +theorem irreducible_units_mul (a : Mˣ) (b : M) : Irreducible (↑a * b) ↔ Irreducible b := by + simp only [irreducible_iff, Units.isUnit_units_mul, and_congr_right_iff] + refine fun _ => ⟨fun h A B HAB => ?_, fun h A B HAB => ?_⟩ + · rw [← a.isUnit_units_mul] + apply h + rw [mul_assoc, ← HAB] + · rw [← a⁻¹.isUnit_units_mul] + apply h + rw [mul_assoc, ← HAB, Units.inv_mul_cancel_left] + +theorem irreducible_isUnit_mul {a b : M} (h : IsUnit a) : Irreducible (a * b) ↔ Irreducible b := + let ⟨a, ha⟩ := h + ha ▸ irreducible_units_mul a b + +theorem irreducible_mul_units (a : Mˣ) (b : M) : Irreducible (b * ↑a) ↔ Irreducible b := by + simp only [irreducible_iff, Units.isUnit_mul_units, and_congr_right_iff] + refine fun _ => ⟨fun h A B HAB => ?_, fun h A B HAB => ?_⟩ + · rw [← Units.isUnit_mul_units B a] + apply h + rw [← mul_assoc, ← HAB] + · rw [← Units.isUnit_mul_units B a⁻¹] + apply h + rw [← mul_assoc, ← HAB, Units.mul_inv_cancel_right] + +theorem irreducible_mul_isUnit {a b : M} (h : IsUnit a) : Irreducible (b * a) ↔ Irreducible b := + let ⟨a, ha⟩ := h + ha ▸ irreducible_mul_units a b + +theorem irreducible_mul_iff {a b : M} : + Irreducible (a * b) ↔ Irreducible a ∧ IsUnit b ∨ Irreducible b ∧ IsUnit a := by + constructor + · refine fun h => Or.imp (fun h' => ⟨?_, h'⟩) (fun h' => ⟨?_, h'⟩) (h.isUnit_or_isUnit rfl).symm + · rwa [irreducible_mul_isUnit h'] at h + · rwa [irreducible_isUnit_mul h'] at h + · rintro (⟨ha, hb⟩ | ⟨hb, ha⟩) + · rwa [irreducible_mul_isUnit hb] + · rwa [irreducible_isUnit_mul ha] + +variable [Monoid N] {F : Type*} [EquivLike F M N] [MulEquivClass F M N] (f : F) + +open MulEquiv + +/-- +Irreducibility is preserved by multiplicative equivalences. +Note that surjective + local hom is not enough. Consider the additive monoids `M = ℕ ⊕ ℕ`, `N = ℕ`, +with a surjective local (additive) hom `f : M →+ N` sending `(m, n)` to `2m + n`. +It is local because the only add unit in `N` is `0`, with preimage `{(0, 0)}` also an add unit. +Then `x = (1, 0)` is irreducible in `M`, but `f x = 2 = 1 + 1` is not irreducible in `N`. +-/ +theorem Irreducible.map {x : M} (h : Irreducible x) : Irreducible (f x) := + ⟨fun g ↦ h.not_unit g.of_map, fun a b g ↦ + let f := MulEquivClass.toMulEquiv f + (h.isUnit_or_isUnit (symm_apply_apply f x ▸ map_mul f.symm a b ▸ congrArg f.symm g)).imp + (·.of_map) (·.of_map)⟩ + +theorem MulEquiv.irreducible_iff (f : F) {a : M} : + Irreducible (f a) ↔ Irreducible a := + ⟨Irreducible.of_map, Irreducible.map f⟩ + +end + +section CommMonoid + +variable [CommMonoid M] {a : M} + +theorem Irreducible.not_square (ha : Irreducible a) : ¬IsSquare a := by + rw [isSquare_iff_exists_sq] + rintro ⟨b, rfl⟩ + exact not_irreducible_pow (by decide) ha + +theorem IsSquare.not_irreducible (ha : IsSquare a) : ¬Irreducible a := fun h => h.not_square ha + +end CommMonoid + +section CancelCommMonoidWithZero + +variable [CancelCommMonoidWithZero M] {a p : M} + +theorem succ_dvd_or_succ_dvd_of_succ_sum_dvd_mul (hp : Prime p) {a b : M} {k l : ℕ} : + p ^ k ∣ a → p ^ l ∣ b → p ^ (k + l + 1) ∣ a * b → p ^ (k + 1) ∣ a ∨ p ^ (l + 1) ∣ b := + fun ⟨x, hx⟩ ⟨y, hy⟩ ⟨z, hz⟩ => + have h : p ^ (k + l) * (x * y) = p ^ (k + l) * (p * z) := by + simpa [mul_comm, pow_add, hx, hy, mul_assoc, mul_left_comm] using hz + have hp0 : p ^ (k + l) ≠ 0 := pow_ne_zero _ hp.ne_zero + have hpd : p ∣ x * y := ⟨z, by rwa [mul_right_inj' hp0] at h⟩ + (hp.dvd_or_dvd hpd).elim + (fun ⟨d, hd⟩ => Or.inl ⟨d, by simp [*, pow_succ, mul_comm, mul_left_comm, mul_assoc]⟩) + fun ⟨d, hd⟩ => Or.inr ⟨d, by simp [*, pow_succ, mul_comm, mul_left_comm, mul_assoc]⟩ + +theorem Prime.not_square (hp : Prime p) : ¬IsSquare p := + hp.irreducible.not_square + +theorem IsSquare.not_prime (ha : IsSquare a) : ¬Prime a := fun h => h.not_square ha + +theorem not_prime_pow {n : ℕ} (hn : n ≠ 1) : ¬Prime (a ^ n) := fun hp => + not_irreducible_pow hn hp.irreducible + +end CancelCommMonoidWithZero + +section CommMonoidWithZero + +theorem DvdNotUnit.isUnit_of_irreducible_right [CommMonoidWithZero M] {p q : M} + (h : DvdNotUnit p q) (hq : Irreducible q) : IsUnit p := by + obtain ⟨_, x, hx, hx'⟩ := h + exact Or.resolve_right ((irreducible_iff.1 hq).right p x hx') hx + +theorem not_irreducible_of_not_unit_dvdNotUnit [CommMonoidWithZero M] {p q : M} (hp : ¬IsUnit p) + (h : DvdNotUnit p q) : ¬Irreducible q := + mt h.isUnit_of_irreducible_right hp + +theorem DvdNotUnit.not_unit [CommMonoidWithZero M] {p q : M} (hp : DvdNotUnit p q) : ¬IsUnit q := by + obtain ⟨-, x, hx, rfl⟩ := hp + exact fun hc => hx (isUnit_iff_dvd_one.mpr (dvd_of_mul_left_dvd (isUnit_iff_dvd_one.mp hc))) + +end CommMonoidWithZero + +section CancelCommMonoidWithZero + +theorem DvdNotUnit.ne [CancelCommMonoidWithZero M] {p q : M} (h : DvdNotUnit p q) : p ≠ q := by + by_contra hcontra + obtain ⟨hp, x, hx', hx''⟩ := h + conv_lhs at hx'' => rw [← hcontra, ← mul_one p] + rw [(mul_left_cancel₀ hp hx'').symm] at hx' + exact hx' isUnit_one + +theorem pow_injective_of_not_isUnit [CancelCommMonoidWithZero M] {q : M} (hq : ¬IsUnit q) + (hq' : q ≠ 0) : Function.Injective fun n : ℕ => q ^ n := by + refine injective_of_lt_imp_ne fun n m h => DvdNotUnit.ne ⟨pow_ne_zero n hq', q ^ (m - n), ?_, ?_⟩ + · exact not_isUnit_of_not_isUnit_dvd hq (dvd_pow (dvd_refl _) (Nat.sub_pos_of_lt h).ne') + · exact (pow_mul_pow_sub q h.le).symm + +@[deprecated (since := "2024-09-22")] +alias pow_injective_of_not_unit := pow_injective_of_not_isUnit + +theorem pow_inj_of_not_isUnit [CancelCommMonoidWithZero M] {q : M} (hq : ¬IsUnit q) + (hq' : q ≠ 0) {m n : ℕ} : q ^ m = q ^ n ↔ m = n := + (pow_injective_of_not_isUnit hq hq').eq_iff + +end CancelCommMonoidWithZero diff --git a/Mathlib/Data/Nat/Prime/Basic.lean b/Mathlib/Data/Nat/Prime/Basic.lean index d1a2155b074ae..5a4b6a4fb3db7 100644 --- a/Mathlib/Data/Nat/Prime/Basic.lean +++ b/Mathlib/Data/Nat/Prime/Basic.lean @@ -3,11 +3,12 @@ Copyright (c) 2015 Microsoft Corporation. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Jeremy Avigad, Mario Carneiro -/ +import Mathlib.Algebra.Associated.Basic +import Mathlib.Algebra.Order.Monoid.Unbundled.Pow import Mathlib.Algebra.Ring.Int -import Mathlib.Order.Bounds.Basic import Mathlib.Data.Nat.Factorial.Basic import Mathlib.Data.Nat.Prime.Defs -import Mathlib.Algebra.Order.Monoid.Unbundled.Pow +import Mathlib.Order.Bounds.Basic /-! ## Notable Theorems diff --git a/Mathlib/Data/Nat/Prime/Defs.lean b/Mathlib/Data/Nat/Prime/Defs.lean index 638f139f80dd6..1741083417a66 100644 --- a/Mathlib/Data/Nat/Prime/Defs.lean +++ b/Mathlib/Data/Nat/Prime/Defs.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Jeremy Avigad, Mario Carneiro -/ import Batteries.Data.Nat.Gcd -import Mathlib.Algebra.Associated.Basic +import Mathlib.Algebra.Prime.Lemmas import Mathlib.Algebra.Ring.Parity /-! diff --git a/Mathlib/RingTheory/Prime.lean b/Mathlib/RingTheory/Prime.lean index 6fb8a1e0598db..331460fc59ec6 100644 --- a/Mathlib/RingTheory/Prime.lean +++ b/Mathlib/RingTheory/Prime.lean @@ -3,10 +3,11 @@ Copyright (c) 2020 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes -/ -import Mathlib.Algebra.Associated.Basic import Mathlib.Algebra.BigOperators.Group.Finset import Mathlib.Algebra.Ring.Divisibility.Basic import Mathlib.Algebra.Order.Group.Unbundled.Abs +import Mathlib.Algebra.Prime.Defs +import Mathlib.Algebra.Ring.Units /-! # Prime elements in rings From a1cf0e3bef13bfbc210c074e07fd7bc9b78852d6 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 25 Oct 2024 13:43:26 +0000 Subject: [PATCH 416/505] feat: Matrices over finite free modules are finite modules (#18098) The titular result is trivial, but to state the lemmas about the basis this uses, we need to generalize many results about `stdBasisMatrix`. --- Mathlib/Data/Matrix/Basis.lean | 53 ++++++++++++------- .../Dimension/Constructions.lean | 40 +++++++++----- .../FreeModule/Finite/Basic.lean | 12 +++-- Mathlib/LinearAlgebra/Matrix/Ideal.lean | 2 +- Mathlib/LinearAlgebra/Matrix/ToLin.lean | 2 +- Mathlib/LinearAlgebra/StdBasis.lean | 20 +++++++ Mathlib/RingTheory/MatrixAlgebra.lean | 6 +-- 7 files changed, 94 insertions(+), 41 deletions(-) diff --git a/Mathlib/Data/Matrix/Basis.lean b/Mathlib/Data/Matrix/Basis.lean index 4b4af6da0eab3..23333213c522e 100644 --- a/Mathlib/Data/Matrix/Basis.lean +++ b/Mathlib/Data/Matrix/Basis.lean @@ -22,13 +22,15 @@ namespace Matrix open Matrix variable [DecidableEq l] [DecidableEq m] [DecidableEq n] -variable [Semiring α] + +section Zero +variable [Zero α] /-- `stdBasisMatrix i j a` is the matrix with `a` in the `i`-th row, `j`-th column, and zeroes elsewhere. -/ -def stdBasisMatrix (i : m) (j : n) (a : α) : Matrix m n α := fun i' j' => - if i = i' ∧ j = j' then a else 0 +def stdBasisMatrix (i : m) (j : n) (a : α) : Matrix m n α := + of <| fun i' j' => if i = i' ∧ j = j' then a else 0 theorem stdBasisMatrix_eq_of_single_single (i : m) (j : n) (a : α) : stdBasisMatrix i j a = Matrix.of (Pi.single i (Pi.single j a)) := by @@ -49,12 +51,16 @@ theorem stdBasisMatrix_zero (i : m) (j : n) : stdBasisMatrix i j (0 : α) = 0 := ext simp -theorem stdBasisMatrix_add (i : m) (j : n) (a b : α) : +end Zero + +theorem stdBasisMatrix_add [AddZeroClass α] (i : m) (j : n) (a b : α) : stdBasisMatrix i j (a + b) = stdBasisMatrix i j a + stdBasisMatrix i j b := by - unfold stdBasisMatrix; ext + ext + simp only [stdBasisMatrix, of_apply] split_ifs with h <;> simp [h] -theorem mulVec_stdBasisMatrix [Fintype m] (i : n) (j : m) (c : α) (x : m → α) : +theorem mulVec_stdBasisMatrix [NonUnitalNonAssocSemiring α] [Fintype m] + (i : n) (j : m) (c : α) (x : m → α) : mulVec (stdBasisMatrix i j c) x = Function.update (0 : n → α) i (c * x j) := by ext i' simp [stdBasisMatrix, mulVec, dotProduct] @@ -62,7 +68,7 @@ theorem mulVec_stdBasisMatrix [Fintype m] (i : n) (j : m) (c : α) (x : m → α · simp simp [h, h.symm] -theorem matrix_eq_sum_stdBasisMatrix [Fintype m] [Fintype n] (x : Matrix m n α) : +theorem matrix_eq_sum_stdBasisMatrix [AddCommMonoid α] [Fintype m] [Fintype n] (x : Matrix m n α) : x = ∑ i : m, ∑ j : n, stdBasisMatrix i j (x i j) := by ext i j; symm iterate 2 rw [Finset.sum_apply] @@ -75,7 +81,7 @@ theorem matrix_eq_sum_stdBasisMatrix [Fintype m] [Fintype n] (x : Matrix m n α) @[deprecated (since := "2024-08-11")] alias matrix_eq_sum_std_basis := matrix_eq_sum_stdBasisMatrix -theorem stdBasisMatrix_eq_single_vecMulVec_single (i : m) (j : n) : +theorem stdBasisMatrix_eq_single_vecMulVec_single [MulZeroOneClass α] (i : m) (j : n) : stdBasisMatrix i j (1 : α) = vecMulVec (Pi.single i 1) (Pi.single j 1) := by ext i' j' -- Porting note: lean3 didn't apply `mul_ite`. @@ -84,7 +90,7 @@ theorem stdBasisMatrix_eq_single_vecMulVec_single (i : m) (j : n) : -- TODO: tie this up with the `Basis` machinery of linear algebra -- this is not completely trivial because we are indexing by two types, instead of one @[deprecated stdBasisMatrix_eq_single_vecMulVec_single (since := "2024-08-11")] -theorem std_basis_eq_basis_mul_basis (i : m) (j : n) : +theorem std_basis_eq_basis_mul_basis [MulZeroOneClass α] (i : m) (j : n) : stdBasisMatrix i j (1 : α) = vecMulVec (fun i' => ite (i = i') 1 0) fun j' => ite (j = j') 1 0 := by rw [stdBasisMatrix_eq_single_vecMulVec_single] @@ -92,7 +98,8 @@ theorem std_basis_eq_basis_mul_basis (i : m) (j : n) : -- todo: the old proof used fintypes, I don't know `Finsupp` but this feels generalizable @[elab_as_elim] -protected theorem induction_on' [Finite m] [Finite n] {P : Matrix m n α → Prop} (M : Matrix m n α) +protected theorem induction_on' + [AddCommMonoid α] [Finite m] [Finite n] {P : Matrix m n α → Prop} (M : Matrix m n α) (h_zero : P 0) (h_add : ∀ p q, P p → P q → P (p + q)) (h_std_basis : ∀ (i : m) (j : n) (x : α), P (stdBasisMatrix i j x)) : P M := by cases nonempty_fintype m; cases nonempty_fintype n @@ -102,7 +109,8 @@ protected theorem induction_on' [Finite m] [Finite n] {P : Matrix m n α → Pro apply h_std_basis @[elab_as_elim] -protected theorem induction_on [Finite m] [Finite n] [Nonempty m] [Nonempty n] +protected theorem induction_on + [AddCommMonoid α] [Finite m] [Finite n] [Nonempty m] [Nonempty n] {P : Matrix m n α → Prop} (M : Matrix m n α) (h_add : ∀ p q, P p → P q → P (p + q)) (h_std_basis : ∀ i j x, P (stdBasisMatrix i j x)) : P M := Matrix.induction_on' M @@ -116,7 +124,7 @@ namespace StdBasisMatrix section -variable (i : m) (j : n) (c : α) (i' : m) (j' : n) +variable [Zero α] (i : m) (j : n) (c : α) (i' : m) (j' : n) @[simp] theorem apply_same : stdBasisMatrix i j c i j = c := @@ -124,7 +132,7 @@ theorem apply_same : stdBasisMatrix i j c i j = c := @[simp] theorem apply_of_ne (h : ¬(i = i' ∧ j = j')) : stdBasisMatrix i j c i' j' = 0 := by - simp only [stdBasisMatrix, and_imp, ite_eq_right_iff] + simp only [stdBasisMatrix, and_imp, ite_eq_right_iff, of_apply] tauto @[simp] @@ -138,8 +146,7 @@ theorem apply_of_col_ne (i i' : m) {j j' : n} (hj : j ≠ j') (a : α) : end section - -variable (i j : n) (c : α) (i' j' : n) +variable [Zero α] (i j : n) (c : α) @[simp] theorem diag_zero (h : j ≠ i) : diag (stdBasisMatrix i j c) = 0 := @@ -150,7 +157,10 @@ theorem diag_same : diag (stdBasisMatrix i i c) = Pi.single i c := by ext j by_cases hij : i = j <;> (try rw [hij]) <;> simp [hij] -variable [Fintype n] +end + +section trace +variable [Fintype n] [AddCommMonoid α] (i j : n) (c : α) @[simp] theorem trace_zero (h : j ≠ i) : trace (stdBasisMatrix i j c) = 0 := by @@ -162,6 +172,11 @@ theorem trace_eq : trace (stdBasisMatrix i i c) = c := by -- Porting note: added `-diag_apply` simp [trace, -diag_apply] +end trace + +section mul +variable [Fintype n] [NonUnitalNonAssocSemiring α] (i j : n) (c : α) + @[simp] theorem mul_left_apply_same (b : n) (M : Matrix n n α) : (stdBasisMatrix i j c * M) i b = c * M j b := by simp [mul_apply, stdBasisMatrix] @@ -189,7 +204,7 @@ theorem mul_same (k : n) (d : α) : theorem mul_of_ne {k l : n} (h : j ≠ k) (d : α) : stdBasisMatrix i j c * stdBasisMatrix k l d = 0 := by ext a b - simp only [mul_apply, boole_mul, stdBasisMatrix] + simp only [mul_apply, boole_mul, stdBasisMatrix, of_apply] by_cases h₁ : i = a -- porting note (#10745): was `simp [h₁, h, h.symm]` · simp only [h₁, true_and, mul_ite, ite_mul, zero_mul, mul_zero, ← ite_and, zero_apply] @@ -200,13 +215,13 @@ theorem mul_of_ne {k l : n} (h : j ≠ k) (d : α) : · simp only [h₁, false_and, ite_false, mul_ite, zero_mul, mul_zero, ite_self, Finset.sum_const_zero, zero_apply] -end +end mul end StdBasisMatrix section Commute -variable [Fintype n] +variable [Fintype n] [Semiring α] theorem row_eq_zero_of_commute_stdBasisMatrix {i j k : n} {M : Matrix n n α} (hM : Commute (stdBasisMatrix i j 1) M) (hkj : k ≠ j) : M j k = 0 := by diff --git a/Mathlib/LinearAlgebra/Dimension/Constructions.lean b/Mathlib/LinearAlgebra/Dimension/Constructions.lean index d03d8646c1451..fb4695c2deb46 100644 --- a/Mathlib/LinearAlgebra/Dimension/Constructions.lean +++ b/Mathlib/LinearAlgebra/Dimension/Constructions.lean @@ -183,25 +183,41 @@ theorem rank_directSum {ι : Type v} (M : ι → Type w) [∀ i : ι, AddCommGro let b : Basis _ R (⨁ i, M i) := DFinsupp.basis fun i => B i simp [← b.mk_eq_rank'', fun i => (B i).mk_eq_rank''] -/-- If `m` and `n` are `Fintype`, the rank of `m × n` matrices is `(#m).lift * (#n).lift`. -/ +/-- If `m` and `n` are finite, the rank of `m × n` matrices over a module `M` is +`(#m).lift * (#n).lift * rank R M`. -/ @[simp] +theorem rank_matrix_module (m : Type w) (n : Type w') [Finite m] [Finite n] : + Module.rank R (Matrix m n M) = + lift.{max v w'} #m * lift.{max v w} #n * lift.{max w w'} (Module.rank R M) := by + cases nonempty_fintype m + cases nonempty_fintype n + obtain ⟨I, b⟩ := Module.Free.exists_basis (R := R) (M := M) + rw [← (b.matrix m n).mk_eq_rank''] + simp only [mk_prod, lift_mul, lift_lift, ← mul_assoc, b.mk_eq_rank''] + + +/-- If `m` and `n` are finite and lie in the same universe, the rank of `m × n` matrices over a +module `M` is `(#m * #n).lift * rank R M`. -/ +@[simp high] +theorem rank_matrix_module' (m n : Type w) [Finite m] [Finite n] : + Module.rank R (Matrix m n M) = + lift.{max v} (#m * #n) * lift.{w} (Module.rank R M) := by + rw [rank_matrix_module, lift_mul, lift_umax.{w, v}] + +/-- If `m` and `n` are finite, the rank of `m × n` matrices is `(#m).lift * (#n).lift`. -/ theorem rank_matrix (m : Type v) (n : Type w) [Finite m] [Finite n] : Module.rank R (Matrix m n R) = Cardinal.lift.{max v w u, v} #m * Cardinal.lift.{max v w u, w} #n := by - cases nonempty_fintype m - cases nonempty_fintype n - have h := (Matrix.stdBasis R m n).mk_eq_rank - rw [← lift_lift.{max v w u, max v w}, lift_inj] at h - simpa using h.symm + rw [rank_matrix_module, rank_self, lift_one, mul_one, ← lift_lift.{v, max u w}, lift_id, + ← lift_lift.{w, max u v}, lift_id] -/-- If `m` and `n` are `Fintype` that lie in the same universe, the rank of `m × n` matrices is +/-- If `m` and `n` are finite and lie in the same universe, the rank of `m × n` matrices is `(#n * #m).lift`. -/ -@[simp high] theorem rank_matrix' (m n : Type v) [Finite m] [Finite n] : Module.rank R (Matrix m n R) = Cardinal.lift.{u} (#m * #n) := by rw [rank_matrix, lift_mul, lift_umax.{v, u}] -/-- If `m` and `n` are `Fintype` that lie in the same universe as `R`, the rank of `m × n` matrices +/-- If `m` and `n` are finite and lie in the same universe as `R`, the rank of `m × n` matrices is `# m * # n`. -/ theorem rank_matrix'' (m n : Type u) [Finite m] [Finite n] : Module.rank R (Matrix m n R) = #m * #n := by simp @@ -228,10 +244,10 @@ theorem finrank_directSum {ι : Type v} [Fintype ι] (M : ι → Type w) [∀ i simp only [finrank, fun i => rank_eq_card_chooseBasisIndex R (M i), rank_directSum, ← mk_sigma, mk_toNat_eq_card, card_sigma] -/-- If `m` and `n` are `Fintype`, the finrank of `m × n` matrices is - `(Fintype.card m) * (Fintype.card n)`. -/ +/-- If `m` and `n` are `Fintype`, the finrank of `m × n` matrices over a module `M` is + `(Fintype.card m) * (Fintype.card n) * finrank R M`. -/ theorem finrank_matrix (m n : Type*) [Fintype m] [Fintype n] : - finrank R (Matrix m n R) = card m * card n := by simp [finrank] + finrank R (Matrix m n M) = card m * card n * finrank R M := by simp [finrank] end Module diff --git a/Mathlib/LinearAlgebra/FreeModule/Finite/Basic.lean b/Mathlib/LinearAlgebra/FreeModule/Finite/Basic.lean index 7cc74f0c2a909..d9da43ef8b1b4 100644 --- a/Mathlib/LinearAlgebra/FreeModule/Finite/Basic.lean +++ b/Mathlib/LinearAlgebra/FreeModule/Finite/Basic.lean @@ -39,9 +39,13 @@ theorem Module.Finite.of_basis {R M ι : Type*} [Semiring R] [AddCommMonoid M] [ refine ⟨⟨Finset.univ.image b, ?_⟩⟩ simp only [Set.image_univ, Finset.coe_univ, Finset.coe_image, Basis.span_eq] -instance Module.Finite.matrix {R : Type u} [Semiring R] - {ι₁ ι₂ : Type*} [_root_.Finite ι₁] [_root_.Finite ι₂] : - Module.Finite R (Matrix ι₁ ι₂ R) := by +instance Module.Finite.matrix {R ι₁ ι₂ M : Type*} + [Semiring R] [AddCommMonoid M] [Module R M] [Module.Free R M] [Module.Finite R M] + [_root_.Finite ι₁] [_root_.Finite ι₂] : + Module.Finite R (Matrix ι₁ ι₂ M) := by cases nonempty_fintype ι₁ cases nonempty_fintype ι₂ - exact Module.Finite.of_basis (Pi.basis fun _ => Pi.basisFun R _) + exact Module.Finite.of_basis <| (Free.chooseBasis _ _).matrix _ _ + +example {ι₁ ι₂ R : Type*} [Semiring R] [Finite ι₁] [Finite ι₂] : + Module.Finite R (Matrix ι₁ ι₂ R) := inferInstance diff --git a/Mathlib/LinearAlgebra/Matrix/Ideal.lean b/Mathlib/LinearAlgebra/Matrix/Ideal.lean index 2d0f2ed989b04..b6372378bf74b 100644 --- a/Mathlib/LinearAlgebra/Matrix/Ideal.lean +++ b/Mathlib/LinearAlgebra/Matrix/Ideal.lean @@ -159,7 +159,7 @@ def equivMatricesOver (i j : n) : TwoSidedIdeal R ≃ TwoSidedIdeal (Matrix n n by_cases hab : a = k ∧ b = l · rcases hab with ⟨ha, hb⟩ subst ha hb - simp only [stdBasisMatrix, and_self, ↓reduceIte, StdBasisMatrix.mul_right_apply_same, + simp only [StdBasisMatrix.apply_same, StdBasisMatrix.mul_right_apply_same, StdBasisMatrix.mul_left_apply_same, one_mul, mul_one] rw [hy2 a b] · conv_lhs => diff --git a/Mathlib/LinearAlgebra/Matrix/ToLin.lean b/Mathlib/LinearAlgebra/Matrix/ToLin.lean index de60812076e72..693066a10400f 100644 --- a/Mathlib/LinearAlgebra/Matrix/ToLin.lean +++ b/Mathlib/LinearAlgebra/Matrix/ToLin.lean @@ -960,7 +960,7 @@ lemma linearMap_apply_apply (ij : ι₂ × ι₁) (k : ι₁) : (b₁.linearMap b₂ ij) (b₁ k) = if ij.2 = k then b₂ ij.1 else 0 := by have := Classical.decEq ι₂ rw [linearMap_apply, Matrix.stdBasis_eq_stdBasisMatrix, Matrix.toLin_self] - dsimp only [Matrix.stdBasisMatrix] + dsimp only [Matrix.stdBasisMatrix, of_apply] simp_rw [ite_smul, one_smul, zero_smul, ite_and, Finset.sum_ite_eq, Finset.mem_univ, if_true] /-- The standard basis of the endomorphism algebra of a module diff --git a/Mathlib/LinearAlgebra/StdBasis.lean b/Mathlib/LinearAlgebra/StdBasis.lean index 3159915328c21..8896d7312a535 100644 --- a/Mathlib/LinearAlgebra/StdBasis.lean +++ b/Mathlib/LinearAlgebra/StdBasis.lean @@ -287,6 +287,26 @@ lemma piEquiv_apply_apply (ι R M : Type*) [Fintype ι] [CommSemiring R] end Module +namespace Basis +variable {ι R M : Type*} (m n : Type*) +variable [Fintype m] [Fintype n] [Semiring R] [AddCommMonoid M] [Module R M] + +/-- The standard basis of `Matrix m n M` given a basis on `M`. -/ +protected noncomputable def matrix (b : Basis ι R M) : + Basis (m × n × ι) R (Matrix m n M) := + Basis.reindex (Pi.basis fun _ : m => Pi.basis fun _ : n => b) + ((Equiv.sigmaEquivProd _ _).trans <| .prodCongr (.refl _) (Equiv.sigmaEquivProd _ _)) + |>.map (Matrix.ofLinearEquiv R) + +variable {n m} + +@[simp] +theorem matrix_apply (b : Basis ι R M) (i : m) (j : n) (k : ι) [DecidableEq m] [DecidableEq n] : + b.matrix m n (i, j, k) = Matrix.stdBasisMatrix i j (b k) := by + simp [Basis.matrix, Matrix.stdBasisMatrix_eq_of_single_single] + +end Basis + namespace Matrix variable (R : Type*) (m n : Type*) [Fintype m] [Finite n] [Semiring R] diff --git a/Mathlib/RingTheory/MatrixAlgebra.lean b/Mathlib/RingTheory/MatrixAlgebra.lean index 4a71967b89d79..0750d23424344 100644 --- a/Mathlib/RingTheory/MatrixAlgebra.lean +++ b/Mathlib/RingTheory/MatrixAlgebra.lean @@ -99,13 +99,11 @@ theorem invFun_algebraMap (M : Matrix n n R) : invFun R A n (M.map (algebraMap R convert Finset.sum_product (β := Matrix n n R) ..; simp theorem right_inv (M : Matrix n n A) : (toFunAlgHom R A n) (invFun R A n M) = M := by - simp only [invFun, map_sum, stdBasisMatrix, apply_ite ↑(algebraMap R A), smul_eq_mul, - mul_boole, toFunAlgHom_apply, RingHom.map_zero, RingHom.map_one, Matrix.map_apply, - Pi.smul_def] + simp only [invFun, map_sum, toFunAlgHom_apply] convert Finset.sum_product (β := Matrix n n A) .. conv_lhs => rw [matrix_eq_sum_stdBasisMatrix M] refine Finset.sum_congr rfl fun i _ => Finset.sum_congr rfl fun j _ => Matrix.ext fun a b => ?_ - simp only [stdBasisMatrix, smul_apply, Matrix.map_apply] + dsimp [stdBasisMatrix] split_ifs <;> aesop theorem left_inv (M : A ⊗[R] Matrix n n R) : invFun R A n (toFunAlgHom R A n M) = M := by From 27e4d0a1b1a1869163126f8e0fdeddd4daf5dbed Mon Sep 17 00:00:00 2001 From: Jiang Jiedong Date: Fri, 25 Oct 2024 14:17:18 +0000 Subject: [PATCH 417/505] feat(Topology/Algebra/Valued/NormedValued): add lemmas for `Valued.toNormedField` (#16758) Add simp lemmas for `Valued.toNormedField`. Prove `isNonarchimedean` for this norm. Add definition `Valued.toNontriviallyNormedField`. --- .../Topology/Algebra/Valued/NormedValued.lean | 50 +++++++++++++++++++ 1 file changed, 50 insertions(+) diff --git a/Mathlib/Topology/Algebra/Valued/NormedValued.lean b/Mathlib/Topology/Algebra/Valued/NormedValued.lean index 249d5250b9868..19c82cd591ba9 100644 --- a/Mathlib/Topology/Algebra/Valued/NormedValued.lean +++ b/Mathlib/Topology/Algebra/Valued/NormedValued.lean @@ -160,4 +160,54 @@ lemma coe_valuation_eq_rankOne_hom_comp_valuation : ⇑NormedField.valuation = h end NormedField +variable {L} {Γ₀} + +namespace toNormedField + +variable {x x' : L} + +@[simp] +theorem norm_le_iff : ‖x‖ ≤ ‖x'‖ ↔ val.v x ≤ val.v x' := + (Valuation.RankOne.strictMono val.v).le_iff_le + +@[simp] +theorem norm_lt_iff : ‖x‖ < ‖x'‖ ↔ val.v x < val.v x' := + (Valuation.RankOne.strictMono val.v).lt_iff_lt + +@[simp] +theorem norm_le_one_iff : ‖x‖ ≤ 1 ↔ val.v x ≤ 1 := by + simpa only [_root_.map_one] using (Valuation.RankOne.strictMono val.v).le_iff_le (b := 1) + +@[simp] +theorem norm_lt_one_iff : ‖x‖ < 1 ↔ val.v x < 1 := by + simpa only [_root_.map_one] using (Valuation.RankOne.strictMono val.v).lt_iff_lt (b := 1) + +@[simp] +theorem one_le_norm_iff : 1 ≤ ‖x‖ ↔ 1 ≤ val.v x := by + simpa only [_root_.map_one] using (Valuation.RankOne.strictMono val.v).le_iff_le (a := 1) + +@[simp] +theorem one_lt_norm_iff : 1 < ‖x‖ ↔ 1 < val.v x := by + simpa only [_root_.map_one] using (Valuation.RankOne.strictMono val.v).lt_iff_lt (a := 1) + +end toNormedField + +/-- +The nontrivially normed field structure determined by a rank one valuation. +-/ +def toNontriviallyNormedField: NontriviallyNormedField L := { + val.toNormedField with + non_trivial := by + obtain ⟨x, hx⟩ := Valuation.RankOne.nontrivial val.v + rcases Valuation.val_le_one_or_val_inv_le_one val.v x with h | h + · use x⁻¹ + simp only [toNormedField.one_lt_norm_iff, map_inv₀, one_lt_inv₀ (zero_lt_iff.mpr hx.1), + lt_of_le_of_ne h hx.2] + · use x + simp only [map_inv₀, inv_le_one₀ <| zero_lt_iff.mpr hx.1] at h + simp only [toNormedField.one_lt_norm_iff, lt_of_le_of_ne h hx.2.symm] +} + +scoped[Valued] attribute [instance] Valued.toNontriviallyNormedField + end Valued From a4411e690f209860046c5e29ee95aac2addf7218 Mon Sep 17 00:00:00 2001 From: damiano Date: Fri, 25 Oct 2024 14:41:59 +0000 Subject: [PATCH 418/505] chore: remove unused variables (#18226) #17715 --- Mathlib/LinearAlgebra/Dimension/Localization.lean | 2 +- Mathlib/LinearAlgebra/Matrix/Circulant.lean | 2 +- Mathlib/LinearAlgebra/RootSystem/OfBilinear.lean | 2 +- Mathlib/MeasureTheory/Decomposition/RadonNikodym.lean | 2 +- Mathlib/MeasureTheory/Decomposition/SignedLebesgue.lean | 2 +- .../Function/ConditionalExpectation/AEMeasurable.lean | 4 +--- .../MeasureTheory/Measure/WithDensityVectorMeasure.lean | 4 ++-- Mathlib/ModelTheory/Graph.lean | 4 ++-- Mathlib/ModelTheory/PartialEquiv.lean | 4 ++-- Mathlib/NumberTheory/Cyclotomic/Gal.lean | 2 +- Mathlib/NumberTheory/Dioph.lean | 4 ++-- Mathlib/Order/Ideal.lean | 8 ++++---- Mathlib/RingTheory/DedekindDomain/Dvr.lean | 2 +- Mathlib/RingTheory/DiscreteValuationRing/TFAE.lean | 2 +- Mathlib/RingTheory/PowerSeries/Basic.lean | 2 +- Mathlib/RingTheory/PowerSeries/Inverse.lean | 2 +- 16 files changed, 23 insertions(+), 25 deletions(-) diff --git a/Mathlib/LinearAlgebra/Dimension/Localization.lean b/Mathlib/LinearAlgebra/Dimension/Localization.lean index b6445fadfe529..594d62bba3e1d 100644 --- a/Mathlib/LinearAlgebra/Dimension/Localization.lean +++ b/Mathlib/LinearAlgebra/Dimension/Localization.lean @@ -120,7 +120,7 @@ end CommRing section Ring -variable {R} [Ring R] [IsDomain R] (S : Submonoid R) +variable {R} [Ring R] [IsDomain R] /-- A domain that is not (left) Ore is of infinite rank. See [cohn_1995] Proposition 1.3.6 -/ diff --git a/Mathlib/LinearAlgebra/Matrix/Circulant.lean b/Mathlib/LinearAlgebra/Matrix/Circulant.lean index c81affc351505..2e87e67a1fa42 100644 --- a/Mathlib/LinearAlgebra/Matrix/Circulant.lean +++ b/Mathlib/LinearAlgebra/Matrix/Circulant.lean @@ -32,7 +32,7 @@ circulant, matrix -/ -variable {α β m n R : Type*} +variable {α β n R : Type*} namespace Matrix diff --git a/Mathlib/LinearAlgebra/RootSystem/OfBilinear.lean b/Mathlib/LinearAlgebra/RootSystem/OfBilinear.lean index b339fa3a131cc..ae27ada2167b0 100644 --- a/Mathlib/LinearAlgebra/RootSystem/OfBilinear.lean +++ b/Mathlib/LinearAlgebra/RootSystem/OfBilinear.lean @@ -25,7 +25,7 @@ open Set Function Module noncomputable section -variable {ι R M N : Type*} [CommRing R] [AddCommGroup M] [Module R M] [AddCommGroup N] [Module R N] +variable {R M : Type*} [CommRing R] [AddCommGroup M] [Module R M] namespace LinearMap diff --git a/Mathlib/MeasureTheory/Decomposition/RadonNikodym.lean b/Mathlib/MeasureTheory/Decomposition/RadonNikodym.lean index 83a31d388715d..133145051b3f2 100644 --- a/Mathlib/MeasureTheory/Decomposition/RadonNikodym.lean +++ b/Mathlib/MeasureTheory/Decomposition/RadonNikodym.lean @@ -88,7 +88,7 @@ lemma rnDeriv_pos' [HaveLebesgueDecomposition ν μ] [SigmaFinite μ] (hμν : section rnDeriv_withDensity_leftRight -variable {μ ν : Measure α} {f : α → ℝ≥0∞} +variable {f : α → ℝ≥0∞} /-- Auxiliary lemma for `rnDeriv_withDensity_left`. -/ lemma rnDeriv_withDensity_withDensity_rnDeriv_left (μ ν : Measure α) [SigmaFinite μ] [SigmaFinite ν] diff --git a/Mathlib/MeasureTheory/Decomposition/SignedLebesgue.lean b/Mathlib/MeasureTheory/Decomposition/SignedLebesgue.lean index a74a832aa9956..7652993309eaf 100644 --- a/Mathlib/MeasureTheory/Decomposition/SignedLebesgue.lean +++ b/Mathlib/MeasureTheory/Decomposition/SignedLebesgue.lean @@ -46,7 +46,7 @@ open scoped Classical MeasureTheory NNReal ENNReal open Set -variable {α β : Type*} {m : MeasurableSpace α} {μ ν : MeasureTheory.Measure α} +variable {α : Type*} {m : MeasurableSpace α} {μ : MeasureTheory.Measure α} namespace MeasureTheory diff --git a/Mathlib/MeasureTheory/Function/ConditionalExpectation/AEMeasurable.lean b/Mathlib/MeasureTheory/Function/ConditionalExpectation/AEMeasurable.lean index bf52c95cfaeaa..95319082eb996 100644 --- a/Mathlib/MeasureTheory/Function/ConditionalExpectation/AEMeasurable.lean +++ b/Mathlib/MeasureTheory/Function/ConditionalExpectation/AEMeasurable.lean @@ -175,12 +175,10 @@ theorem AEStronglyMeasurable'.aeStronglyMeasurable'_of_measurableSpace_le_on {α hf_ind.stronglyMeasurable_of_measurableSpace_le_on hs_m hs fun x hxs => Set.indicator_of_not_mem hxs _ -variable {α F F' 𝕜 : Type*} {p : ℝ≥0∞} [RCLike 𝕜] +variable {α F 𝕜 : Type*} {p : ℝ≥0∞} [RCLike 𝕜] -- 𝕜 for ℝ or ℂ -- F for a Lp submodule [NormedAddCommGroup F] [NormedSpace 𝕜 F] - -- F' for integrals on a Lp submodule - [NormedAddCommGroup F'] [NormedSpace 𝕜 F'] [NormedSpace ℝ F'] section LpMeas diff --git a/Mathlib/MeasureTheory/Measure/WithDensityVectorMeasure.lean b/Mathlib/MeasureTheory/Measure/WithDensityVectorMeasure.lean index 674e34c9ca6f3..095d95dd608a0 100644 --- a/Mathlib/MeasureTheory/Measure/WithDensityVectorMeasure.lean +++ b/Mathlib/MeasureTheory/Measure/WithDensityVectorMeasure.lean @@ -26,13 +26,13 @@ noncomputable section open scoped MeasureTheory NNReal ENNReal -variable {α β : Type*} {m : MeasurableSpace α} +variable {α : Type*} {m : MeasurableSpace α} namespace MeasureTheory open TopologicalSpace -variable {μ ν : Measure α} +variable {μ : Measure α} variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] open Classical in diff --git a/Mathlib/ModelTheory/Graph.lean b/Mathlib/ModelTheory/Graph.lean index c4ce433290e9c..ed3a509c17c55 100644 --- a/Mathlib/ModelTheory/Graph.lean +++ b/Mathlib/ModelTheory/Graph.lean @@ -21,7 +21,7 @@ This file defines first-order languages, structures, and theories in graph theor of the theory of simple graphs. -/ -universe u v +universe u namespace FirstOrder @@ -31,7 +31,7 @@ open FirstOrder open Structure -variable {α : Type u} {V : Type v} {n : ℕ} +variable {V : Type u} {n : ℕ} /-! ### Simple Graphs -/ diff --git a/Mathlib/ModelTheory/PartialEquiv.lean b/Mathlib/ModelTheory/PartialEquiv.lean index 3db2f33308397..49e73c8f359f2 100644 --- a/Mathlib/ModelTheory/PartialEquiv.lean +++ b/Mathlib/ModelTheory/PartialEquiv.lean @@ -40,8 +40,8 @@ namespace FirstOrder namespace Language -variable (L : Language.{u, v}) (M : Type w) (N : Type w') {P : Type*} -variable [L.Structure M] [L.Structure N] [L.Structure P] +variable (L : Language.{u, v}) (M : Type w) (N : Type w') +variable [L.Structure M] [L.Structure N] open FirstOrder Structure Substructure diff --git a/Mathlib/NumberTheory/Cyclotomic/Gal.lean b/Mathlib/NumberTheory/Cyclotomic/Gal.lean index fb381459d36cf..c5f27c52696f2 100644 --- a/Mathlib/NumberTheory/Cyclotomic/Gal.lean +++ b/Mathlib/NumberTheory/Cyclotomic/Gal.lean @@ -152,7 +152,7 @@ end IsCyclotomicExtension section Gal -variable [Field L] (hμ : IsPrimitiveRoot μ n) [Algebra K L] [IsCyclotomicExtension {n} K L] +variable [Field L] [Algebra K L] [IsCyclotomicExtension {n} K L] (h : Irreducible (cyclotomic n K)) {K} /-- `IsCyclotomicExtension.autEquivPow` repackaged in terms of `Gal`. diff --git a/Mathlib/NumberTheory/Dioph.lean b/Mathlib/NumberTheory/Dioph.lean index 721b81b5271ba..6c7c378d55eb4 100644 --- a/Mathlib/NumberTheory/Dioph.lean +++ b/Mathlib/NumberTheory/Dioph.lean @@ -67,7 +67,7 @@ Note that this duplicates `MvPolynomial`. section Polynomials -variable {α β γ : Type*} +variable {α β : Type*} /-- A predicate asserting that a function is a multivariate integer polynomial. (We are being a bit lazy here by allowing many representations for multiplication, @@ -432,7 +432,7 @@ end section -variable {α β : Type} {n : ℕ} +variable {α : Type} {n : ℕ} open Vector3 diff --git a/Mathlib/Order/Ideal.lean b/Mathlib/Order/Ideal.lean index 4dae232039074..7c8fe0e4ef08f 100644 --- a/Mathlib/Order/Ideal.lean +++ b/Mathlib/Order/Ideal.lean @@ -92,7 +92,7 @@ variable [LE P] section -variable {I J s t : Ideal P} {x y : P} +variable {I s t : Ideal P} {x : P} theorem toLowerSet_injective : Injective (toLowerSet : Ideal P → LowerSet P) := fun s t _ ↦ by cases s @@ -247,7 +247,7 @@ variable [Preorder P] section -variable {I J : Ideal P} {x y : P} +variable {I : Ideal P} {x y : P} /-- The smallest ideal containing a given element. -/ @[simps] @@ -316,7 +316,7 @@ end SemilatticeSup section SemilatticeSupDirected -variable [SemilatticeSup P] [IsDirected P (· ≥ ·)] {x : P} {I J K s t : Ideal P} +variable [SemilatticeSup P] [IsDirected P (· ≥ ·)] {x : P} {I J s t : Ideal P} /-- The infimum of two ideals of a co-directed order is their intersection. -/ instance : Inf (Ideal P) := @@ -385,7 +385,7 @@ end SemilatticeSupDirected section SemilatticeSupOrderBot -variable [SemilatticeSup P] [OrderBot P] {x : P} {I J K : Ideal P} +variable [SemilatticeSup P] [OrderBot P] {x : P} instance : InfSet (Ideal P) := ⟨fun S ↦ diff --git a/Mathlib/RingTheory/DedekindDomain/Dvr.lean b/Mathlib/RingTheory/DedekindDomain/Dvr.lean index 2edc66bfc7bfa..c7086a988650b 100644 --- a/Mathlib/RingTheory/DedekindDomain/Dvr.lean +++ b/Mathlib/RingTheory/DedekindDomain/Dvr.lean @@ -43,7 +43,7 @@ dedekind domain, dedekind ring -/ -variable (R A K : Type*) [CommRing R] [CommRing A] [IsDomain A] [Field K] +variable (A : Type*) [CommRing A] [IsDomain A] open scoped nonZeroDivisors Polynomial diff --git a/Mathlib/RingTheory/DiscreteValuationRing/TFAE.lean b/Mathlib/RingTheory/DiscreteValuationRing/TFAE.lean index 5957f090f0b89..763ee85bc69d5 100644 --- a/Mathlib/RingTheory/DiscreteValuationRing/TFAE.lean +++ b/Mathlib/RingTheory/DiscreteValuationRing/TFAE.lean @@ -25,7 +25,7 @@ Also see `tfae_of_isNoetherianRing_of_localRing_of_isDomain` for a version witho -/ -variable (R : Type*) [CommRing R] (K : Type*) [Field K] [Algebra R K] [IsFractionRing R K] +variable (R : Type*) [CommRing R] open scoped Multiplicative diff --git a/Mathlib/RingTheory/PowerSeries/Basic.lean b/Mathlib/RingTheory/PowerSeries/Basic.lean index 65b589c7373de..8965ee6ddf242 100644 --- a/Mathlib/RingTheory/PowerSeries/Basic.lean +++ b/Mathlib/RingTheory/PowerSeries/Basic.lean @@ -758,7 +758,7 @@ namespace Polynomial open Finsupp Polynomial -variable {σ : Type*} {R : Type*} [CommSemiring R] (φ ψ : R[X]) +variable {R : Type*} [CommSemiring R] (φ ψ : R[X]) -- Porting note: added so we can add the `@[coe]` attribute /-- The natural inclusion from polynomials into formal power series. -/ diff --git a/Mathlib/RingTheory/PowerSeries/Inverse.lean b/Mathlib/RingTheory/PowerSeries/Inverse.lean index 18c6913c525cd..660923cc4da32 100644 --- a/Mathlib/RingTheory/PowerSeries/Inverse.lean +++ b/Mathlib/RingTheory/PowerSeries/Inverse.lean @@ -274,7 +274,7 @@ theorem map.isLocalHom : IsLocalHom (map f) := @[deprecated (since := "2024-10-10")] alias map.isLocalRingHom := map.isLocalHom -variable [LocalRing R] [LocalRing S] +variable [LocalRing R] instance : LocalRing R⟦X⟧ := { inferInstanceAs <| LocalRing <| MvPowerSeries Unit R with } From aea6f296a35d3b88812a1b424543e1da53822850 Mon Sep 17 00:00:00 2001 From: Markus Himmel Date: Fri, 25 Oct 2024 14:58:57 +0000 Subject: [PATCH 419/505] chore: split `Limits/Shapes/Terminal.lean` to reduce imports (#18229) The parts that require `HasLimit` stay in `Terminal.lean`, and the parts that only require `IsLimit` go to `IsTerminal.lean`. Found using the new `lake exe unused`. Co-authored-by: Markus Himmel --- Mathlib.lean | 1 + Mathlib/CategoryTheory/Adjunction/Comma.lean | 1 + .../CategoryTheory/Bicategory/Kan/HasKan.lean | 1 + Mathlib/CategoryTheory/Comma/Over.lean | 4 +- .../Comma/StructuredArrow/Basic.lean | 2 +- Mathlib/CategoryTheory/Elements.lean | 3 +- .../CategoryTheory/Endofunctor/Algebra.lean | 2 +- .../Limits/Shapes/IsTerminal.lean | 438 ++++++++++++++++++ .../Limits/Shapes/Terminal.lean | 384 +-------------- Mathlib/CategoryTheory/Monad/Limits.lean | 2 +- Mathlib/CategoryTheory/WithTerminal.lean | 2 +- 11 files changed, 448 insertions(+), 392 deletions(-) create mode 100644 Mathlib/CategoryTheory/Limits/Shapes/IsTerminal.lean diff --git a/Mathlib.lean b/Mathlib.lean index d885b3af7da93..05b32bb4e3391 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1733,6 +1733,7 @@ import Mathlib.CategoryTheory.Limits.Shapes.FiniteProducts import Mathlib.CategoryTheory.Limits.Shapes.FunctorCategory import Mathlib.CategoryTheory.Limits.Shapes.FunctorToTypes import Mathlib.CategoryTheory.Limits.Shapes.Images +import Mathlib.CategoryTheory.Limits.Shapes.IsTerminal import Mathlib.CategoryTheory.Limits.Shapes.KernelPair import Mathlib.CategoryTheory.Limits.Shapes.Kernels import Mathlib.CategoryTheory.Limits.Shapes.Multiequalizer diff --git a/Mathlib/CategoryTheory/Adjunction/Comma.lean b/Mathlib/CategoryTheory/Adjunction/Comma.lean index 6f4474b55e705..4147a489b5b31 100644 --- a/Mathlib/CategoryTheory/Adjunction/Comma.lean +++ b/Mathlib/CategoryTheory/Adjunction/Comma.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Bhavik Mehta. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Bhavik Mehta -/ +import Mathlib.CategoryTheory.Limits.Shapes.Terminal import Mathlib.CategoryTheory.Adjunction.Basic import Mathlib.CategoryTheory.Comma.StructuredArrow.Basic import Mathlib.CategoryTheory.PUnit diff --git a/Mathlib/CategoryTheory/Bicategory/Kan/HasKan.lean b/Mathlib/CategoryTheory/Bicategory/Kan/HasKan.lean index e81755e4324af..40043a409ecfe 100644 --- a/Mathlib/CategoryTheory/Bicategory/Kan/HasKan.lean +++ b/Mathlib/CategoryTheory/Bicategory/Kan/HasKan.lean @@ -3,6 +3,7 @@ Copyright (c) 2024 Yuma Mizuno. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yuma Mizuno -/ +import Mathlib.CategoryTheory.Limits.Shapes.Terminal import Mathlib.CategoryTheory.Bicategory.Kan.IsKan /-! diff --git a/Mathlib/CategoryTheory/Comma/Over.lean b/Mathlib/CategoryTheory/Comma/Over.lean index 2d0ed2bd61b9c..e37bf7ec99866 100644 --- a/Mathlib/CategoryTheory/Comma/Over.lean +++ b/Mathlib/CategoryTheory/Comma/Over.lean @@ -4,9 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin, Bhavik Mehta -/ import Mathlib.CategoryTheory.Comma.StructuredArrow.Basic -import Mathlib.CategoryTheory.PUnit -import Mathlib.CategoryTheory.Functor.ReflectsIso -import Mathlib.CategoryTheory.Functor.EpiMono +import Mathlib.CategoryTheory.Category.Cat /-! # Over and under categories diff --git a/Mathlib/CategoryTheory/Comma/StructuredArrow/Basic.lean b/Mathlib/CategoryTheory/Comma/StructuredArrow/Basic.lean index c0d1729054be6..7e4231ecc0c06 100644 --- a/Mathlib/CategoryTheory/Comma/StructuredArrow/Basic.lean +++ b/Mathlib/CategoryTheory/Comma/StructuredArrow/Basic.lean @@ -5,7 +5,7 @@ Authors: Adam Topaz, Kim Morrison -/ import Mathlib.CategoryTheory.Comma.Basic import Mathlib.CategoryTheory.PUnit -import Mathlib.CategoryTheory.Limits.Shapes.Terminal +import Mathlib.CategoryTheory.Limits.Shapes.IsTerminal /-! # The category of "structured arrows" diff --git a/Mathlib/CategoryTheory/Elements.lean b/Mathlib/CategoryTheory/Elements.lean index 2c682ecb881f9..7d85416b30755 100644 --- a/Mathlib/CategoryTheory/Elements.lean +++ b/Mathlib/CategoryTheory/Elements.lean @@ -4,8 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison -/ import Mathlib.CategoryTheory.Comma.StructuredArrow.Basic -import Mathlib.CategoryTheory.Groupoid -import Mathlib.CategoryTheory.PUnit +import Mathlib.CategoryTheory.Category.Cat /-! # The category of elements diff --git a/Mathlib/CategoryTheory/Endofunctor/Algebra.lean b/Mathlib/CategoryTheory/Endofunctor/Algebra.lean index 59d61ad43c4c1..1b53aa148eb3a 100644 --- a/Mathlib/CategoryTheory/Endofunctor/Algebra.lean +++ b/Mathlib/CategoryTheory/Endofunctor/Algebra.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Joseph Hua. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison, Bhavik Mehta, Johan Commelin, Reid Barton, Robert Y. Lewis, Joseph Hua -/ -import Mathlib.CategoryTheory.Limits.Shapes.Terminal +import Mathlib.CategoryTheory.Limits.Shapes.IsTerminal /-! diff --git a/Mathlib/CategoryTheory/Limits/Shapes/IsTerminal.lean b/Mathlib/CategoryTheory/Limits/Shapes/IsTerminal.lean new file mode 100644 index 0000000000000..7ceb4a6aeb551 --- /dev/null +++ b/Mathlib/CategoryTheory/Limits/Shapes/IsTerminal.lean @@ -0,0 +1,438 @@ +/- +Copyright (c) 2019 Kim Morrison. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kim Morrison, Bhavik Mehta +-/ +import Mathlib.CategoryTheory.PEmpty +import Mathlib.CategoryTheory.Limits.IsLimit +import Mathlib.CategoryTheory.EpiMono +import Mathlib.CategoryTheory.Category.Preorder + +/-! +# Initial and terminal objects in a category. + +In this file we define the predicates `IsTerminal` and `IsInitial` as well as the class +`InitialMonoClass`. + +The classes `HasTerminal` and `HasInitial` and the associated notations for terminal and inital +objects are defined in `Terminal.lean`. + +## References +* [Stacks: Initial and final objects](https://stacks.math.columbia.edu/tag/002B) +-/ + +assert_not_exists CategoryTheory.Limits.HasLimit + +noncomputable section + +universe w w' v v₁ v₂ u u₁ u₂ + +open CategoryTheory + +namespace CategoryTheory.Limits + +variable {C : Type u₁} [Category.{v₁} C] + +/-- Construct a cone for the empty diagram given an object. -/ +@[simps] +def asEmptyCone (X : C) : Cone (Functor.empty.{0} C) := + { pt := X + π := + { app := by aesop_cat } } + +/-- Construct a cocone for the empty diagram given an object. -/ +@[simps] +def asEmptyCocone (X : C) : Cocone (Functor.empty.{0} C) := + { pt := X + ι := + { app := by aesop_cat } } + +/-- `X` is terminal if the cone it induces on the empty diagram is limiting. -/ +abbrev IsTerminal (X : C) := + IsLimit (asEmptyCone X) + +/-- `X` is initial if the cocone it induces on the empty diagram is colimiting. -/ +abbrev IsInitial (X : C) := + IsColimit (asEmptyCocone X) + +/-- An object `Y` is terminal iff for every `X` there is a unique morphism `X ⟶ Y`. -/ +def isTerminalEquivUnique (F : Discrete.{0} PEmpty.{1} ⥤ C) (Y : C) : + IsLimit (⟨Y, by aesop_cat, by aesop_cat⟩ : Cone F) ≃ ∀ X : C, Unique (X ⟶ Y) where + toFun t X := + { default := t.lift ⟨X, ⟨by aesop_cat, by aesop_cat⟩⟩ + uniq := fun f => + t.uniq ⟨X, ⟨by aesop_cat, by aesop_cat⟩⟩ f (by aesop_cat) } + invFun u := + { lift := fun s => (u s.pt).default + uniq := fun s _ _ => (u s.pt).2 _ } + left_inv := by dsimp [Function.LeftInverse]; intro x; simp only [eq_iff_true_of_subsingleton] + right_inv := by + dsimp [Function.RightInverse,Function.LeftInverse] + intro u; funext X; simp only + +/-- An object `Y` is terminal if for every `X` there is a unique morphism `X ⟶ Y` + (as an instance). -/ +def IsTerminal.ofUnique (Y : C) [h : ∀ X : C, Unique (X ⟶ Y)] : IsTerminal Y where + lift s := (h s.pt).default + fac := fun _ ⟨j⟩ => j.elim + +/-- An object `Y` is terminal if for every `X` there is a unique morphism `X ⟶ Y` + (as explicit arguments). -/ +def IsTerminal.ofUniqueHom {Y : C} (h : ∀ X : C, X ⟶ Y) (uniq : ∀ (X : C) (m : X ⟶ Y), m = h X) : + IsTerminal Y := + have : ∀ X : C, Unique (X ⟶ Y) := fun X ↦ ⟨⟨h X⟩, uniq X⟩ + IsTerminal.ofUnique Y + +/-- If `α` is a preorder with top, then `⊤` is a terminal object. -/ +def isTerminalTop {α : Type*} [Preorder α] [OrderTop α] : IsTerminal (⊤ : α) := + IsTerminal.ofUnique _ + +/-- Transport a term of type `IsTerminal` across an isomorphism. -/ +def IsTerminal.ofIso {Y Z : C} (hY : IsTerminal Y) (i : Y ≅ Z) : IsTerminal Z := + IsLimit.ofIsoLimit hY + { hom := { hom := i.hom } + inv := { hom := i.inv } } + +/-- If `X` and `Y` are isomorphic, then `X` is terminal iff `Y` is. -/ +def IsTerminal.equivOfIso {X Y : C} (e : X ≅ Y) : + IsTerminal X ≃ IsTerminal Y where + toFun h := IsTerminal.ofIso h e + invFun h := IsTerminal.ofIso h e.symm + left_inv _ := Subsingleton.elim _ _ + right_inv _ := Subsingleton.elim _ _ + +/-- An object `X` is initial iff for every `Y` there is a unique morphism `X ⟶ Y`. -/ +def isInitialEquivUnique (F : Discrete.{0} PEmpty.{1} ⥤ C) (X : C) : + IsColimit (⟨X, ⟨by aesop_cat, by aesop_cat⟩⟩ : Cocone F) ≃ ∀ Y : C, Unique (X ⟶ Y) where + toFun t X := + { default := t.desc ⟨X, ⟨by aesop_cat, by aesop_cat⟩⟩ + uniq := fun f => t.uniq ⟨X, ⟨by aesop_cat, by aesop_cat⟩⟩ f (by aesop_cat) } + invFun u := + { desc := fun s => (u s.pt).default + uniq := fun s _ _ => (u s.pt).2 _ } + left_inv := by dsimp [Function.LeftInverse]; intro; simp only [eq_iff_true_of_subsingleton] + right_inv := by + dsimp [Function.RightInverse,Function.LeftInverse] + intro; funext; simp only + +/-- An object `X` is initial if for every `Y` there is a unique morphism `X ⟶ Y` + (as an instance). -/ +def IsInitial.ofUnique (X : C) [h : ∀ Y : C, Unique (X ⟶ Y)] : IsInitial X where + desc s := (h s.pt).default + fac := fun _ ⟨j⟩ => j.elim + +/-- An object `X` is initial if for every `Y` there is a unique morphism `X ⟶ Y` + (as explicit arguments). -/ +def IsInitial.ofUniqueHom {X : C} (h : ∀ Y : C, X ⟶ Y) (uniq : ∀ (Y : C) (m : X ⟶ Y), m = h Y) : + IsInitial X := + have : ∀ Y : C, Unique (X ⟶ Y) := fun Y ↦ ⟨⟨h Y⟩, uniq Y⟩ + IsInitial.ofUnique X + +/-- If `α` is a preorder with bot, then `⊥` is an initial object. -/ +def isInitialBot {α : Type*} [Preorder α] [OrderBot α] : IsInitial (⊥ : α) := + IsInitial.ofUnique _ + +/-- Transport a term of type `is_initial` across an isomorphism. -/ +def IsInitial.ofIso {X Y : C} (hX : IsInitial X) (i : X ≅ Y) : IsInitial Y := + IsColimit.ofIsoColimit hX + { hom := { hom := i.hom } + inv := { hom := i.inv } } + +/-- If `X` and `Y` are isomorphic, then `X` is initial iff `Y` is. -/ +def IsInitial.equivOfIso {X Y : C} (e : X ≅ Y) : + IsInitial X ≃ IsInitial Y where + toFun h := IsInitial.ofIso h e + invFun h := IsInitial.ofIso h e.symm + left_inv _ := Subsingleton.elim _ _ + right_inv _ := Subsingleton.elim _ _ + +/-- Give the morphism to a terminal object from any other. -/ +def IsTerminal.from {X : C} (t : IsTerminal X) (Y : C) : Y ⟶ X := + t.lift (asEmptyCone Y) + +/-- Any two morphisms to a terminal object are equal. -/ +theorem IsTerminal.hom_ext {X Y : C} (t : IsTerminal X) (f g : Y ⟶ X) : f = g := + IsLimit.hom_ext t (by aesop_cat) + +@[simp] +theorem IsTerminal.comp_from {Z : C} (t : IsTerminal Z) {X Y : C} (f : X ⟶ Y) : + f ≫ t.from Y = t.from X := + t.hom_ext _ _ + +@[simp] +theorem IsTerminal.from_self {X : C} (t : IsTerminal X) : t.from X = 𝟙 X := + t.hom_ext _ _ + +/-- Give the morphism from an initial object to any other. -/ +def IsInitial.to {X : C} (t : IsInitial X) (Y : C) : X ⟶ Y := + t.desc (asEmptyCocone Y) + +/-- Any two morphisms from an initial object are equal. -/ +theorem IsInitial.hom_ext {X Y : C} (t : IsInitial X) (f g : X ⟶ Y) : f = g := + IsColimit.hom_ext t (by aesop_cat) + +@[simp] +theorem IsInitial.to_comp {X : C} (t : IsInitial X) {Y Z : C} (f : Y ⟶ Z) : t.to Y ≫ f = t.to Z := + t.hom_ext _ _ + +@[simp] +theorem IsInitial.to_self {X : C} (t : IsInitial X) : t.to X = 𝟙 X := + t.hom_ext _ _ + +/-- Any morphism from a terminal object is split mono. -/ +theorem IsTerminal.isSplitMono_from {X Y : C} (t : IsTerminal X) (f : X ⟶ Y) : IsSplitMono f := + IsSplitMono.mk' ⟨t.from _, t.hom_ext _ _⟩ + +/-- Any morphism to an initial object is split epi. -/ +theorem IsInitial.isSplitEpi_to {X Y : C} (t : IsInitial X) (f : Y ⟶ X) : IsSplitEpi f := + IsSplitEpi.mk' ⟨t.to _, t.hom_ext _ _⟩ + +/-- Any morphism from a terminal object is mono. -/ +theorem IsTerminal.mono_from {X Y : C} (t : IsTerminal X) (f : X ⟶ Y) : Mono f := by + haveI := t.isSplitMono_from f; infer_instance + +/-- Any morphism to an initial object is epi. -/ +theorem IsInitial.epi_to {X Y : C} (t : IsInitial X) (f : Y ⟶ X) : Epi f := by + haveI := t.isSplitEpi_to f; infer_instance + +/-- If `T` and `T'` are terminal, they are isomorphic. -/ +@[simps] +def IsTerminal.uniqueUpToIso {T T' : C} (hT : IsTerminal T) (hT' : IsTerminal T') : T ≅ T' where + hom := hT'.from _ + inv := hT.from _ + +/-- If `I` and `I'` are initial, they are isomorphic. -/ +@[simps] +def IsInitial.uniqueUpToIso {I I' : C} (hI : IsInitial I) (hI' : IsInitial I') : I ≅ I' where + hom := hI.to _ + inv := hI'.to _ + +variable (C) + +section Univ + +variable (X : C) {F₁ : Discrete.{w} PEmpty ⥤ C} {F₂ : Discrete.{w'} PEmpty ⥤ C} + +/-- Being terminal is independent of the empty diagram, its universe, and the cone over it, + as long as the cone points are isomorphic. -/ +def isLimitChangeEmptyCone {c₁ : Cone F₁} (hl : IsLimit c₁) (c₂ : Cone F₂) (hi : c₁.pt ≅ c₂.pt) : + IsLimit c₂ where + lift c := hl.lift ⟨c.pt, by aesop_cat, by aesop_cat⟩ ≫ hi.hom + uniq c f _ := by + dsimp + rw [← hl.uniq _ (f ≫ hi.inv) _] + · simp only [Category.assoc, Iso.inv_hom_id, Category.comp_id] + · aesop_cat + +/-- Replacing an empty cone in `IsLimit` by another with the same cone point + is an equivalence. -/ +def isLimitEmptyConeEquiv (c₁ : Cone F₁) (c₂ : Cone F₂) (h : c₁.pt ≅ c₂.pt) : + IsLimit c₁ ≃ IsLimit c₂ where + toFun hl := isLimitChangeEmptyCone C hl c₂ h + invFun hl := isLimitChangeEmptyCone C hl c₁ h.symm + left_inv := by dsimp [Function.LeftInverse]; intro; simp only [eq_iff_true_of_subsingleton] + right_inv := by + dsimp [Function.LeftInverse,Function.RightInverse]; intro + simp only [eq_iff_true_of_subsingleton] + +/-- Being initial is independent of the empty diagram, its universe, and the cocone over it, + as long as the cocone points are isomorphic. -/ +def isColimitChangeEmptyCocone {c₁ : Cocone F₁} (hl : IsColimit c₁) (c₂ : Cocone F₂) + (hi : c₁.pt ≅ c₂.pt) : IsColimit c₂ where + desc c := hi.inv ≫ hl.desc ⟨c.pt, by aesop_cat, by aesop_cat⟩ + uniq c f _ := by + dsimp + rw [← hl.uniq _ (hi.hom ≫ f) _] + · simp only [Iso.inv_hom_id_assoc] + · aesop_cat + +/-- Replacing an empty cocone in `IsColimit` by another with the same cocone point + is an equivalence. -/ +def isColimitEmptyCoconeEquiv (c₁ : Cocone F₁) (c₂ : Cocone F₂) (h : c₁.pt ≅ c₂.pt) : + IsColimit c₁ ≃ IsColimit c₂ where + toFun hl := isColimitChangeEmptyCocone C hl c₂ h + invFun hl := isColimitChangeEmptyCocone C hl c₁ h.symm + left_inv := by dsimp [Function.LeftInverse]; intro; simp only [eq_iff_true_of_subsingleton] + right_inv := by + dsimp [Function.LeftInverse,Function.RightInverse]; intro + simp only [eq_iff_true_of_subsingleton] + +end Univ + +section + +variable {C} + +/-- An initial object is terminal in the opposite category. -/ +def terminalOpOfInitial {X : C} (t : IsInitial X) : IsTerminal (Opposite.op X) where + lift s := (t.to s.pt.unop).op + uniq _ _ _ := Quiver.Hom.unop_inj (t.hom_ext _ _) + +/-- An initial object in the opposite category is terminal in the original category. -/ +def terminalUnopOfInitial {X : Cᵒᵖ} (t : IsInitial X) : IsTerminal X.unop where + lift s := (t.to (Opposite.op s.pt)).unop + uniq _ _ _ := Quiver.Hom.op_inj (t.hom_ext _ _) + +/-- A terminal object is initial in the opposite category. -/ +def initialOpOfTerminal {X : C} (t : IsTerminal X) : IsInitial (Opposite.op X) where + desc s := (t.from s.pt.unop).op + uniq _ _ _ := Quiver.Hom.unop_inj (t.hom_ext _ _) + +/-- A terminal object in the opposite category is initial in the original category. -/ +def initialUnopOfTerminal {X : Cᵒᵖ} (t : IsTerminal X) : IsInitial X.unop where + desc s := (t.from (Opposite.op s.pt)).unop + uniq _ _ _ := Quiver.Hom.op_inj (t.hom_ext _ _) + +/-- A category is an `InitialMonoClass` if the canonical morphism of an initial object is a +monomorphism. In practice, this is most useful when given an arbitrary morphism out of the chosen +initial object, see `initial.mono_from`. +Given a terminal object, this is equivalent to the assumption that the unique morphism from initial +to terminal is a monomorphism, which is the second of Freyd's axioms for an AT category. + +TODO: This is a condition satisfied by categories with zero objects and morphisms. +-/ +class InitialMonoClass (C : Type u₁) [Category.{v₁} C] : Prop where + /-- The map from the (any as stated) initial object to any other object is a + monomorphism -/ + isInitial_mono_from : ∀ {I} (X : C) (hI : IsInitial I), Mono (hI.to X) + +theorem IsInitial.mono_from [InitialMonoClass C] {I} {X : C} (hI : IsInitial I) (f : I ⟶ X) : + Mono f := by + rw [hI.hom_ext f (hI.to X)] + apply InitialMonoClass.isInitial_mono_from + +/-- To show a category is an `InitialMonoClass` it suffices to give an initial object such that +every morphism out of it is a monomorphism. -/ +theorem InitialMonoClass.of_isInitial {I : C} (hI : IsInitial I) (h : ∀ X, Mono (hI.to X)) : + InitialMonoClass C where + isInitial_mono_from {I'} X hI' := by + rw [hI'.hom_ext (hI'.to X) ((hI'.uniqueUpToIso hI).hom ≫ hI.to X)] + apply mono_comp + +/-- To show a category is an `InitialMonoClass` it suffices to show the unique morphism from an +initial object to a terminal object is a monomorphism. -/ +theorem InitialMonoClass.of_isTerminal {I T : C} (hI : IsInitial I) (hT : IsTerminal T) + (_ : Mono (hI.to T)) : InitialMonoClass C := + InitialMonoClass.of_isInitial hI fun X => mono_of_mono_fac (hI.hom_ext (_ ≫ hT.from X) (hI.to T)) + +section Comparison + +variable {D : Type u₂} [Category.{v₂} D] (G : C ⥤ D) + +end Comparison + +variable {J : Type u} [Category.{v} J] + +/-- From a functor `F : J ⥤ C`, given an initial object of `J`, construct a cone for `J`. +In `limitOfDiagramInitial` we show it is a limit cone. -/ +@[simps] +def coneOfDiagramInitial {X : J} (tX : IsInitial X) (F : J ⥤ C) : Cone F where + pt := F.obj X + π := + { app := fun j => F.map (tX.to j) + naturality := fun j j' k => by + dsimp + rw [← F.map_comp, Category.id_comp, tX.hom_ext (tX.to j ≫ k) (tX.to j')] } + +/-- From a functor `F : J ⥤ C`, given an initial object of `J`, show the cone +`coneOfDiagramInitial` is a limit. -/ +def limitOfDiagramInitial {X : J} (tX : IsInitial X) (F : J ⥤ C) : + IsLimit (coneOfDiagramInitial tX F) where + lift s := s.π.app X + uniq s m w := by + conv_lhs => dsimp + simp_rw [← w X, coneOfDiagramInitial_π_app, tX.hom_ext (tX.to X) (𝟙 _)] + simp + +/-- From a functor `F : J ⥤ C`, given a terminal object of `J`, construct a cone for `J`, +provided that the morphisms in the diagram are isomorphisms. +In `limitOfDiagramTerminal` we show it is a limit cone. -/ +@[simps] +def coneOfDiagramTerminal {X : J} (hX : IsTerminal X) (F : J ⥤ C) + [∀ (i j : J) (f : i ⟶ j), IsIso (F.map f)] : Cone F where + pt := F.obj X + π := + { app := fun _ => inv (F.map (hX.from _)) + naturality := by + intro i j f + dsimp + simp only [IsIso.eq_inv_comp, IsIso.comp_inv_eq, Category.id_comp, ← F.map_comp, + hX.hom_ext (hX.from i) (f ≫ hX.from j)] } + +/-- From a functor `F : J ⥤ C`, given a terminal object of `J` and that the morphisms in the +diagram are isomorphisms, show the cone `coneOfDiagramTerminal` is a limit. -/ +def limitOfDiagramTerminal {X : J} (hX : IsTerminal X) (F : J ⥤ C) + [∀ (i j : J) (f : i ⟶ j), IsIso (F.map f)] : IsLimit (coneOfDiagramTerminal hX F) where + lift S := S.π.app _ + +/-- From a functor `F : J ⥤ C`, given a terminal object of `J`, construct a cocone for `J`. +In `colimitOfDiagramTerminal` we show it is a colimit cocone. -/ +@[simps] +def coconeOfDiagramTerminal {X : J} (tX : IsTerminal X) (F : J ⥤ C) : Cocone F where + pt := F.obj X + ι := + { app := fun j => F.map (tX.from j) + naturality := fun j j' k => by + dsimp + rw [← F.map_comp, Category.comp_id, tX.hom_ext (k ≫ tX.from j') (tX.from j)] } + +/-- From a functor `F : J ⥤ C`, given a terminal object of `J`, show the cocone +`coconeOfDiagramTerminal` is a colimit. -/ +def colimitOfDiagramTerminal {X : J} (tX : IsTerminal X) (F : J ⥤ C) : + IsColimit (coconeOfDiagramTerminal tX F) where + desc s := s.ι.app X + uniq s m w := by + conv_rhs => dsimp -- Porting note: why do I need this much firepower? + rw [← w X, coconeOfDiagramTerminal_ι_app, tX.hom_ext (tX.from X) (𝟙 _)] + simp + +lemma IsColimit.isIso_ι_app_of_isTerminal {F : J ⥤ C} {c : Cocone F} (hc : IsColimit c) + (X : J) (hX : IsTerminal X) : + IsIso (c.ι.app X) := by + change IsIso (coconePointUniqueUpToIso (colimitOfDiagramTerminal hX F) hc).hom + infer_instance + +/-- From a functor `F : J ⥤ C`, given an initial object of `J`, construct a cocone for `J`, +provided that the morphisms in the diagram are isomorphisms. +In `colimitOfDiagramInitial` we show it is a colimit cocone. -/ +@[simps] +def coconeOfDiagramInitial {X : J} (hX : IsInitial X) (F : J ⥤ C) + [∀ (i j : J) (f : i ⟶ j), IsIso (F.map f)] : Cocone F where + pt := F.obj X + ι := + { app := fun _ => inv (F.map (hX.to _)) + naturality := by + intro i j f + dsimp + simp only [IsIso.eq_inv_comp, IsIso.comp_inv_eq, Category.comp_id, ← F.map_comp, + hX.hom_ext (hX.to i ≫ f) (hX.to j)] } + +/-- From a functor `F : J ⥤ C`, given an initial object of `J` and that the morphisms in the +diagram are isomorphisms, show the cone `coconeOfDiagramInitial` is a colimit. -/ +def colimitOfDiagramInitial {X : J} (hX : IsInitial X) (F : J ⥤ C) + [∀ (i j : J) (f : i ⟶ j), IsIso (F.map f)] : IsColimit (coconeOfDiagramInitial hX F) where + desc S := S.ι.app _ + +lemma IsLimit.isIso_π_app_of_isInitial {F : J ⥤ C} {c : Cone F} (hc : IsLimit c) + (X : J) (hX : IsInitial X) : + IsIso (c.π.app X) := by + change IsIso (conePointUniqueUpToIso hc (limitOfDiagramInitial hX F)).hom + infer_instance + +/-- Any morphism between terminal objects is an isomorphism. -/ +lemma isIso_of_isTerminal {X Y : C} (hX : IsTerminal X) (hY : IsTerminal Y) (f : X ⟶ Y) : + IsIso f := by + refine ⟨⟨IsTerminal.from hX Y, ?_⟩⟩ + simp only [IsTerminal.comp_from, IsTerminal.from_self, true_and] + apply IsTerminal.hom_ext hY + +/-- Any morphism between initial objects is an isomorphism. -/ +lemma isIso_of_isInitial {X Y : C} (hX : IsInitial X) (hY : IsInitial Y) (f : X ⟶ Y) : + IsIso f := by + refine ⟨⟨IsInitial.to hY X, ?_⟩⟩ + simp only [IsInitial.to_comp, IsInitial.to_self, and_true] + apply IsInitial.hom_ext hX + +end + +end CategoryTheory.Limits diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Terminal.lean b/Mathlib/CategoryTheory/Limits/Shapes/Terminal.lean index 03741e73c82af..d24c05e73cd2c 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Terminal.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Terminal.lean @@ -3,10 +3,8 @@ Copyright (c) 2019 Kim Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison, Bhavik Mehta -/ -import Mathlib.CategoryTheory.PEmpty +import Mathlib.CategoryTheory.Limits.Shapes.IsTerminal import Mathlib.CategoryTheory.Limits.HasLimits -import Mathlib.CategoryTheory.EpiMono -import Mathlib.CategoryTheory.Category.Preorder /-! # Initial and terminal objects in a category. @@ -26,180 +24,6 @@ namespace CategoryTheory.Limits variable {C : Type u₁} [Category.{v₁} C] -/-- Construct a cone for the empty diagram given an object. -/ -@[simps] -def asEmptyCone (X : C) : Cone (Functor.empty.{0} C) := - { pt := X - π := - { app := by aesop_cat } } - -/-- Construct a cocone for the empty diagram given an object. -/ -@[simps] -def asEmptyCocone (X : C) : Cocone (Functor.empty.{0} C) := - { pt := X - ι := - { app := by aesop_cat } } - -/-- `X` is terminal if the cone it induces on the empty diagram is limiting. -/ -abbrev IsTerminal (X : C) := - IsLimit (asEmptyCone X) - -/-- `X` is initial if the cocone it induces on the empty diagram is colimiting. -/ -abbrev IsInitial (X : C) := - IsColimit (asEmptyCocone X) - -/-- An object `Y` is terminal iff for every `X` there is a unique morphism `X ⟶ Y`. -/ -def isTerminalEquivUnique (F : Discrete.{0} PEmpty.{1} ⥤ C) (Y : C) : - IsLimit (⟨Y, by aesop_cat, by aesop_cat⟩ : Cone F) ≃ ∀ X : C, Unique (X ⟶ Y) where - toFun t X := - { default := t.lift ⟨X, ⟨by aesop_cat, by aesop_cat⟩⟩ - uniq := fun f => - t.uniq ⟨X, ⟨by aesop_cat, by aesop_cat⟩⟩ f (by aesop_cat) } - invFun u := - { lift := fun s => (u s.pt).default - uniq := fun s _ _ => (u s.pt).2 _ } - left_inv := by dsimp [Function.LeftInverse]; intro x; simp only [eq_iff_true_of_subsingleton] - right_inv := by - dsimp [Function.RightInverse,Function.LeftInverse] - intro u; funext X; simp only - -/-- An object `Y` is terminal if for every `X` there is a unique morphism `X ⟶ Y` - (as an instance). -/ -def IsTerminal.ofUnique (Y : C) [h : ∀ X : C, Unique (X ⟶ Y)] : IsTerminal Y where - lift s := (h s.pt).default - fac := fun _ ⟨j⟩ => j.elim - -/-- An object `Y` is terminal if for every `X` there is a unique morphism `X ⟶ Y` - (as explicit arguments). -/ -def IsTerminal.ofUniqueHom {Y : C} (h : ∀ X : C, X ⟶ Y) (uniq : ∀ (X : C) (m : X ⟶ Y), m = h X) : - IsTerminal Y := - have : ∀ X : C, Unique (X ⟶ Y) := fun X ↦ ⟨⟨h X⟩, uniq X⟩ - IsTerminal.ofUnique Y - -/-- If `α` is a preorder with top, then `⊤` is a terminal object. -/ -def isTerminalTop {α : Type*} [Preorder α] [OrderTop α] : IsTerminal (⊤ : α) := - IsTerminal.ofUnique _ - -/-- Transport a term of type `IsTerminal` across an isomorphism. -/ -def IsTerminal.ofIso {Y Z : C} (hY : IsTerminal Y) (i : Y ≅ Z) : IsTerminal Z := - IsLimit.ofIsoLimit hY - { hom := { hom := i.hom } - inv := { hom := i.inv } } - -/-- If `X` and `Y` are isomorphic, then `X` is terminal iff `Y` is. -/ -def IsTerminal.equivOfIso {X Y : C} (e : X ≅ Y) : - IsTerminal X ≃ IsTerminal Y where - toFun h := IsTerminal.ofIso h e - invFun h := IsTerminal.ofIso h e.symm - left_inv _ := Subsingleton.elim _ _ - right_inv _ := Subsingleton.elim _ _ - -/-- An object `X` is initial iff for every `Y` there is a unique morphism `X ⟶ Y`. -/ -def isInitialEquivUnique (F : Discrete.{0} PEmpty.{1} ⥤ C) (X : C) : - IsColimit (⟨X, ⟨by aesop_cat, by aesop_cat⟩⟩ : Cocone F) ≃ ∀ Y : C, Unique (X ⟶ Y) where - toFun t X := - { default := t.desc ⟨X, ⟨by aesop_cat, by aesop_cat⟩⟩ - uniq := fun f => t.uniq ⟨X, ⟨by aesop_cat, by aesop_cat⟩⟩ f (by aesop_cat) } - invFun u := - { desc := fun s => (u s.pt).default - uniq := fun s _ _ => (u s.pt).2 _ } - left_inv := by dsimp [Function.LeftInverse]; intro; simp only [eq_iff_true_of_subsingleton] - right_inv := by - dsimp [Function.RightInverse,Function.LeftInverse] - intro; funext; simp only - -/-- An object `X` is initial if for every `Y` there is a unique morphism `X ⟶ Y` - (as an instance). -/ -def IsInitial.ofUnique (X : C) [h : ∀ Y : C, Unique (X ⟶ Y)] : IsInitial X where - desc s := (h s.pt).default - fac := fun _ ⟨j⟩ => j.elim - -/-- An object `X` is initial if for every `Y` there is a unique morphism `X ⟶ Y` - (as explicit arguments). -/ -def IsInitial.ofUniqueHom {X : C} (h : ∀ Y : C, X ⟶ Y) (uniq : ∀ (Y : C) (m : X ⟶ Y), m = h Y) : - IsInitial X := - have : ∀ Y : C, Unique (X ⟶ Y) := fun Y ↦ ⟨⟨h Y⟩, uniq Y⟩ - IsInitial.ofUnique X - -/-- If `α` is a preorder with bot, then `⊥` is an initial object. -/ -def isInitialBot {α : Type*} [Preorder α] [OrderBot α] : IsInitial (⊥ : α) := - IsInitial.ofUnique _ - -/-- Transport a term of type `is_initial` across an isomorphism. -/ -def IsInitial.ofIso {X Y : C} (hX : IsInitial X) (i : X ≅ Y) : IsInitial Y := - IsColimit.ofIsoColimit hX - { hom := { hom := i.hom } - inv := { hom := i.inv } } - -/-- If `X` and `Y` are isomorphic, then `X` is initial iff `Y` is. -/ -def IsInitial.equivOfIso {X Y : C} (e : X ≅ Y) : - IsInitial X ≃ IsInitial Y where - toFun h := IsInitial.ofIso h e - invFun h := IsInitial.ofIso h e.symm - left_inv _ := Subsingleton.elim _ _ - right_inv _ := Subsingleton.elim _ _ - -/-- Give the morphism to a terminal object from any other. -/ -def IsTerminal.from {X : C} (t : IsTerminal X) (Y : C) : Y ⟶ X := - t.lift (asEmptyCone Y) - -/-- Any two morphisms to a terminal object are equal. -/ -theorem IsTerminal.hom_ext {X Y : C} (t : IsTerminal X) (f g : Y ⟶ X) : f = g := - IsLimit.hom_ext t (by aesop_cat) - -@[simp] -theorem IsTerminal.comp_from {Z : C} (t : IsTerminal Z) {X Y : C} (f : X ⟶ Y) : - f ≫ t.from Y = t.from X := - t.hom_ext _ _ - -@[simp] -theorem IsTerminal.from_self {X : C} (t : IsTerminal X) : t.from X = 𝟙 X := - t.hom_ext _ _ - -/-- Give the morphism from an initial object to any other. -/ -def IsInitial.to {X : C} (t : IsInitial X) (Y : C) : X ⟶ Y := - t.desc (asEmptyCocone Y) - -/-- Any two morphisms from an initial object are equal. -/ -theorem IsInitial.hom_ext {X Y : C} (t : IsInitial X) (f g : X ⟶ Y) : f = g := - IsColimit.hom_ext t (by aesop_cat) - -@[simp] -theorem IsInitial.to_comp {X : C} (t : IsInitial X) {Y Z : C} (f : Y ⟶ Z) : t.to Y ≫ f = t.to Z := - t.hom_ext _ _ - -@[simp] -theorem IsInitial.to_self {X : C} (t : IsInitial X) : t.to X = 𝟙 X := - t.hom_ext _ _ - -/-- Any morphism from a terminal object is split mono. -/ -theorem IsTerminal.isSplitMono_from {X Y : C} (t : IsTerminal X) (f : X ⟶ Y) : IsSplitMono f := - IsSplitMono.mk' ⟨t.from _, t.hom_ext _ _⟩ - -/-- Any morphism to an initial object is split epi. -/ -theorem IsInitial.isSplitEpi_to {X Y : C} (t : IsInitial X) (f : Y ⟶ X) : IsSplitEpi f := - IsSplitEpi.mk' ⟨t.to _, t.hom_ext _ _⟩ - -/-- Any morphism from a terminal object is mono. -/ -theorem IsTerminal.mono_from {X Y : C} (t : IsTerminal X) (f : X ⟶ Y) : Mono f := by - haveI := t.isSplitMono_from f; infer_instance - -/-- Any morphism to an initial object is epi. -/ -theorem IsInitial.epi_to {X Y : C} (t : IsInitial X) (f : Y ⟶ X) : Epi f := by - haveI := t.isSplitEpi_to f; infer_instance - -/-- If `T` and `T'` are terminal, they are isomorphic. -/ -@[simps] -def IsTerminal.uniqueUpToIso {T T' : C} (hT : IsTerminal T) (hT' : IsTerminal T') : T ≅ T' where - hom := hT'.from _ - inv := hT.from _ - -/-- If `I` and `I'` are initial, they are isomorphic. -/ -@[simps] -def IsInitial.uniqueUpToIso {I I' : C} (hI : IsInitial I) (hI' : IsInitial I') : I ≅ I' where - hom := hI.to _ - inv := hI'.to _ - variable (C) /-- A category has a terminal object if it has a limit over the empty diagram. @@ -218,28 +42,6 @@ section Univ variable (X : C) {F₁ : Discrete.{w} PEmpty ⥤ C} {F₂ : Discrete.{w'} PEmpty ⥤ C} -/-- Being terminal is independent of the empty diagram, its universe, and the cone over it, - as long as the cone points are isomorphic. -/ -def isLimitChangeEmptyCone {c₁ : Cone F₁} (hl : IsLimit c₁) (c₂ : Cone F₂) (hi : c₁.pt ≅ c₂.pt) : - IsLimit c₂ where - lift c := hl.lift ⟨c.pt, by aesop_cat, by aesop_cat⟩ ≫ hi.hom - uniq c f _ := by - dsimp - rw [← hl.uniq _ (f ≫ hi.inv) _] - · simp only [Category.assoc, Iso.inv_hom_id, Category.comp_id] - · aesop_cat - -/-- Replacing an empty cone in `IsLimit` by another with the same cone point - is an equivalence. -/ -def isLimitEmptyConeEquiv (c₁ : Cone F₁) (c₂ : Cone F₂) (h : c₁.pt ≅ c₂.pt) : - IsLimit c₁ ≃ IsLimit c₂ where - toFun hl := isLimitChangeEmptyCone C hl c₂ h - invFun hl := isLimitChangeEmptyCone C hl c₁ h.symm - left_inv := by dsimp [Function.LeftInverse]; intro; simp only [eq_iff_true_of_subsingleton] - right_inv := by - dsimp [Function.LeftInverse,Function.RightInverse]; intro - simp only [eq_iff_true_of_subsingleton] - theorem hasTerminalChangeDiagram (h : HasLimit F₁) : HasLimit F₂ := ⟨⟨⟨⟨limit F₁, by aesop_cat, by aesop_cat⟩, isLimitChangeEmptyCone C (limit.isLimit F₁) _ (eqToIso rfl)⟩⟩⟩ @@ -248,28 +50,6 @@ theorem hasTerminalChangeUniverse [h : HasLimitsOfShape (Discrete.{w} PEmpty) C] HasLimitsOfShape (Discrete.{w'} PEmpty) C where has_limit _ := hasTerminalChangeDiagram C (h.1 (Functor.empty C)) -/-- Being initial is independent of the empty diagram, its universe, and the cocone over it, - as long as the cocone points are isomorphic. -/ -def isColimitChangeEmptyCocone {c₁ : Cocone F₁} (hl : IsColimit c₁) (c₂ : Cocone F₂) - (hi : c₁.pt ≅ c₂.pt) : IsColimit c₂ where - desc c := hi.inv ≫ hl.desc ⟨c.pt, by aesop_cat, by aesop_cat⟩ - uniq c f _ := by - dsimp - rw [← hl.uniq _ (hi.hom ≫ f) _] - · simp only [Iso.inv_hom_id_assoc] - · aesop_cat - -/-- Replacing an empty cocone in `IsColimit` by another with the same cocone point - is an equivalence. -/ -def isColimitEmptyCoconeEquiv (c₁ : Cocone F₁) (c₂ : Cocone F₂) (h : c₁.pt ≅ c₂.pt) : - IsColimit c₁ ≃ IsColimit c₂ where - toFun hl := isColimitChangeEmptyCocone C hl c₂ h - invFun hl := isColimitChangeEmptyCocone C hl c₁ h.symm - left_inv := by dsimp [Function.LeftInverse]; intro; simp only [eq_iff_true_of_subsingleton] - right_inv := by - dsimp [Function.LeftInverse,Function.RightInverse]; intro - simp only [eq_iff_true_of_subsingleton] - theorem hasInitialChangeDiagram (h : HasColimit F₁) : HasColimit F₂ := ⟨⟨⟨⟨colimit F₁, by aesop_cat, by aesop_cat⟩, isColimitChangeEmptyCocone C (colimit.isColimit F₁) _ (eqToIso rfl)⟩⟩⟩ @@ -379,26 +159,6 @@ instance terminal.isSplitMono_from {Y : C} [HasTerminal C] (f : ⊤_ C ⟶ Y) : instance initial.isSplitEpi_to {Y : C} [HasInitial C] (f : Y ⟶ ⊥_ C) : IsSplitEpi f := IsInitial.isSplitEpi_to initialIsInitial _ -/-- An initial object is terminal in the opposite category. -/ -def terminalOpOfInitial {X : C} (t : IsInitial X) : IsTerminal (Opposite.op X) where - lift s := (t.to s.pt.unop).op - uniq _ _ _ := Quiver.Hom.unop_inj (t.hom_ext _ _) - -/-- An initial object in the opposite category is terminal in the original category. -/ -def terminalUnopOfInitial {X : Cᵒᵖ} (t : IsInitial X) : IsTerminal X.unop where - lift s := (t.to (Opposite.op s.pt)).unop - uniq _ _ _ := Quiver.Hom.op_inj (t.hom_ext _ _) - -/-- A terminal object is initial in the opposite category. -/ -def initialOpOfTerminal {X : C} (t : IsTerminal X) : IsInitial (Opposite.op X) where - desc s := (t.from s.pt.unop).op - uniq _ _ _ := Quiver.Hom.unop_inj (t.hom_ext _ _) - -/-- A terminal object in the opposite category is initial in the original category. -/ -def initialUnopOfTerminal {X : Cᵒᵖ} (t : IsTerminal X) : IsInitial X.unop where - desc s := (t.from (Opposite.op s.pt)).unop - uniq _ _ _ := Quiver.Hom.op_inj (t.hom_ext _ _) - instance hasInitial_op_of_hasTerminal [HasTerminal C] : HasInitial Cᵒᵖ := (initialOpOfTerminal terminalIsTerminal).hasInitial @@ -459,48 +219,16 @@ theorem ι_colimitConstInitial_hom {J : Type*} [Category J] {C : Type*} [Categor colimit.ι ((CategoryTheory.Functor.const J).obj (⊥_ C)) j ≫ colimitConstInitial.hom = initial.to _ := by aesop_cat -/-- A category is an `InitialMonoClass` if the canonical morphism of an initial object is a -monomorphism. In practice, this is most useful when given an arbitrary morphism out of the chosen -initial object, see `initial.mono_from`. -Given a terminal object, this is equivalent to the assumption that the unique morphism from initial -to terminal is a monomorphism, which is the second of Freyd's axioms for an AT category. - -TODO: This is a condition satisfied by categories with zero objects and morphisms. --/ -class InitialMonoClass (C : Type u₁) [Category.{v₁} C] : Prop where - /-- The map from the (any as stated) initial object to any other object is a - monomorphism -/ - isInitial_mono_from : ∀ {I} (X : C) (hI : IsInitial I), Mono (hI.to X) - -theorem IsInitial.mono_from [InitialMonoClass C] {I} {X : C} (hI : IsInitial I) (f : I ⟶ X) : - Mono f := by - rw [hI.hom_ext f (hI.to X)] - apply InitialMonoClass.isInitial_mono_from - instance (priority := 100) initial.mono_from [HasInitial C] [InitialMonoClass C] (X : C) (f : ⊥_ C ⟶ X) : Mono f := initialIsInitial.mono_from f -/-- To show a category is an `InitialMonoClass` it suffices to give an initial object such that -every morphism out of it is a monomorphism. -/ -theorem InitialMonoClass.of_isInitial {I : C} (hI : IsInitial I) (h : ∀ X, Mono (hI.to X)) : - InitialMonoClass C where - isInitial_mono_from {I'} X hI' := by - rw [hI'.hom_ext (hI'.to X) ((hI'.uniqueUpToIso hI).hom ≫ hI.to X)] - apply mono_comp - /-- To show a category is an `InitialMonoClass` it suffices to show every morphism out of the initial object is a monomorphism. -/ theorem InitialMonoClass.of_initial [HasInitial C] (h : ∀ X : C, Mono (initial.to X)) : InitialMonoClass C := InitialMonoClass.of_isInitial initialIsInitial h -/-- To show a category is an `InitialMonoClass` it suffices to show the unique morphism from an -initial object to a terminal object is a monomorphism. -/ -theorem InitialMonoClass.of_isTerminal {I T : C} (hI : IsInitial I) (hT : IsTerminal T) - (_ : Mono (hI.to T)) : InitialMonoClass C := - InitialMonoClass.of_isInitial hI fun X => mono_of_mono_fac (hI.hom_ext (_ ≫ hT.from X) (hI.to T)) - /-- To show a category is an `InitialMonoClass` it suffices to show the unique morphism from the initial object to a terminal object is a monomorphism. -/ theorem InitialMonoClass.of_terminal [HasInitial C] [HasTerminal C] (h : Mono (initial.to (⊤_ C))) : @@ -531,27 +259,6 @@ end Comparison variable {J : Type u} [Category.{v} J] -/-- From a functor `F : J ⥤ C`, given an initial object of `J`, construct a cone for `J`. -In `limitOfDiagramInitial` we show it is a limit cone. -/ -@[simps] -def coneOfDiagramInitial {X : J} (tX : IsInitial X) (F : J ⥤ C) : Cone F where - pt := F.obj X - π := - { app := fun j => F.map (tX.to j) - naturality := fun j j' k => by - dsimp - rw [← F.map_comp, Category.id_comp, tX.hom_ext (tX.to j ≫ k) (tX.to j')] } - -/-- From a functor `F : J ⥤ C`, given an initial object of `J`, show the cone -`coneOfDiagramInitial` is a limit. -/ -def limitOfDiagramInitial {X : J} (tX : IsInitial X) (F : J ⥤ C) : - IsLimit (coneOfDiagramInitial tX F) where - lift s := s.π.app X - uniq s m w := by - conv_lhs => dsimp - simp_rw [← w X, coneOfDiagramInitial_π_app, tX.hom_ext (tX.to X) (𝟙 _)] - simp - instance hasLimit_of_domain_hasInitial [HasInitial J] {F : J ⥤ C} : HasLimit F := HasLimit.mk { cone := _, isLimit := limitOfDiagramInitial (initialIsInitial) F } @@ -562,27 +269,6 @@ to the limit of `F`. -/ abbrev limitOfInitial (F : J ⥤ C) [HasInitial J] : limit F ≅ F.obj (⊥_ J) := IsLimit.conePointUniqueUpToIso (limit.isLimit _) (limitOfDiagramInitial initialIsInitial F) -/-- From a functor `F : J ⥤ C`, given a terminal object of `J`, construct a cone for `J`, -provided that the morphisms in the diagram are isomorphisms. -In `limitOfDiagramTerminal` we show it is a limit cone. -/ -@[simps] -def coneOfDiagramTerminal {X : J} (hX : IsTerminal X) (F : J ⥤ C) - [∀ (i j : J) (f : i ⟶ j), IsIso (F.map f)] : Cone F where - pt := F.obj X - π := - { app := fun _ => inv (F.map (hX.from _)) - naturality := by - intro i j f - dsimp - simp only [IsIso.eq_inv_comp, IsIso.comp_inv_eq, Category.id_comp, ← F.map_comp, - hX.hom_ext (hX.from i) (f ≫ hX.from j)] } - -/-- From a functor `F : J ⥤ C`, given a terminal object of `J` and that the morphisms in the -diagram are isomorphisms, show the cone `coneOfDiagramTerminal` is a limit. -/ -def limitOfDiagramTerminal {X : J} (hX : IsTerminal X) (F : J ⥤ C) - [∀ (i j : J) (f : i ⟶ j), IsIso (F.map f)] : IsLimit (coneOfDiagramTerminal hX F) where - lift S := S.π.app _ - instance hasLimit_of_domain_hasTerminal [HasTerminal J] {F : J ⥤ C} [∀ (i j : J) (f : i ⟶ j), IsIso (F.map f)] : HasLimit F := HasLimit.mk { cone := _, isLimit := limitOfDiagramTerminal (terminalIsTerminal) F } @@ -594,36 +280,9 @@ abbrev limitOfTerminal (F : J ⥤ C) [HasTerminal J] [∀ (i j : J) (f : i ⟶ j limit F ≅ F.obj (⊤_ J) := IsLimit.conePointUniqueUpToIso (limit.isLimit _) (limitOfDiagramTerminal terminalIsTerminal F) -/-- From a functor `F : J ⥤ C`, given a terminal object of `J`, construct a cocone for `J`. -In `colimitOfDiagramTerminal` we show it is a colimit cocone. -/ -@[simps] -def coconeOfDiagramTerminal {X : J} (tX : IsTerminal X) (F : J ⥤ C) : Cocone F where - pt := F.obj X - ι := - { app := fun j => F.map (tX.from j) - naturality := fun j j' k => by - dsimp - rw [← F.map_comp, Category.comp_id, tX.hom_ext (k ≫ tX.from j') (tX.from j)] } - -/-- From a functor `F : J ⥤ C`, given a terminal object of `J`, show the cocone -`coconeOfDiagramTerminal` is a colimit. -/ -def colimitOfDiagramTerminal {X : J} (tX : IsTerminal X) (F : J ⥤ C) : - IsColimit (coconeOfDiagramTerminal tX F) where - desc s := s.ι.app X - uniq s m w := by - conv_rhs => dsimp -- Porting note: why do I need this much firepower? - rw [← w X, coconeOfDiagramTerminal_ι_app, tX.hom_ext (tX.from X) (𝟙 _)] - simp - instance hasColimit_of_domain_hasTerminal [HasTerminal J] {F : J ⥤ C} : HasColimit F := HasColimit.mk { cocone := _, isColimit := colimitOfDiagramTerminal (terminalIsTerminal) F } -lemma IsColimit.isIso_ι_app_of_isTerminal {F : J ⥤ C} {c : Cocone F} (hc : IsColimit c) - (X : J) (hX : IsTerminal X) : - IsIso (c.ι.app X) := by - change IsIso (coconePointUniqueUpToIso (colimitOfDiagramTerminal hX F) hc).hom - infer_instance - -- This is reducible to allow usage of lemmas about `cocone_point_unique_up_to_iso`. /-- For a functor `F : J ⥤ C`, if `J` has a terminal object then the image of it is isomorphic to the colimit of `F`. -/ @@ -631,37 +290,10 @@ abbrev colimitOfTerminal (F : J ⥤ C) [HasTerminal J] : colimit F ≅ F.obj ( IsColimit.coconePointUniqueUpToIso (colimit.isColimit _) (colimitOfDiagramTerminal terminalIsTerminal F) -/-- From a functor `F : J ⥤ C`, given an initial object of `J`, construct a cocone for `J`, -provided that the morphisms in the diagram are isomorphisms. -In `colimitOfDiagramInitial` we show it is a colimit cocone. -/ -@[simps] -def coconeOfDiagramInitial {X : J} (hX : IsInitial X) (F : J ⥤ C) - [∀ (i j : J) (f : i ⟶ j), IsIso (F.map f)] : Cocone F where - pt := F.obj X - ι := - { app := fun _ => inv (F.map (hX.to _)) - naturality := by - intro i j f - dsimp - simp only [IsIso.eq_inv_comp, IsIso.comp_inv_eq, Category.comp_id, ← F.map_comp, - hX.hom_ext (hX.to i ≫ f) (hX.to j)] } - -/-- From a functor `F : J ⥤ C`, given an initial object of `J` and that the morphisms in the -diagram are isomorphisms, show the cone `coconeOfDiagramInitial` is a colimit. -/ -def colimitOfDiagramInitial {X : J} (hX : IsInitial X) (F : J ⥤ C) - [∀ (i j : J) (f : i ⟶ j), IsIso (F.map f)] : IsColimit (coconeOfDiagramInitial hX F) where - desc S := S.ι.app _ - instance hasColimit_of_domain_hasInitial [HasInitial J] {F : J ⥤ C} [∀ (i j : J) (f : i ⟶ j), IsIso (F.map f)] : HasColimit F := HasColimit.mk { cocone := _, isColimit := colimitOfDiagramInitial (initialIsInitial) F } -lemma IsLimit.isIso_π_app_of_isInitial {F : J ⥤ C} {c : Cone F} (hc : IsLimit c) - (X : J) (hX : IsInitial X) : - IsIso (c.π.app X) := by - change IsIso (conePointUniqueUpToIso hc (limitOfDiagramInitial hX F)).hom - infer_instance - -- This is reducible to allow usage of lemmas about `cocone_point_unique_up_to_iso`. /-- For a functor `F : J ⥤ C`, if `J` has an initial object and all the morphisms in the diagram are isomorphisms, then the image of the initial object is isomorphic to the colimit of `F`. -/ @@ -710,20 +342,6 @@ instance isIso_ι_initial [HasInitial J] (F : J ⥤ C) [∀ (i j : J) (f : i ⟶ IsIso (colimit.ι F (⊥_ J)) := isIso_ι_of_isInitial initialIsInitial F -/-- Any morphism between terminal objects is an isomorphism. -/ -lemma isIso_of_isTerminal {X Y : C} (hX : IsTerminal X) (hY : IsTerminal Y) (f : X ⟶ Y) : - IsIso f := by - refine ⟨⟨IsTerminal.from hX Y, ?_⟩⟩ - simp only [IsTerminal.comp_from, IsTerminal.from_self, true_and] - apply IsTerminal.hom_ext hY - -/-- Any morphism between initial objects is an isomorphism. -/ -lemma isIso_of_isInitial {X Y : C} (hX : IsInitial X) (hY : IsInitial Y) (f : X ⟶ Y) : - IsIso f := by - refine ⟨⟨IsInitial.to hY X, ?_⟩⟩ - simp only [IsInitial.to_comp, IsInitial.to_self, and_true] - apply IsInitial.hom_ext hX - end end CategoryTheory.Limits diff --git a/Mathlib/CategoryTheory/Monad/Limits.lean b/Mathlib/CategoryTheory/Monad/Limits.lean index 1e1128d105453..11d531a830a9b 100644 --- a/Mathlib/CategoryTheory/Monad/Limits.lean +++ b/Mathlib/CategoryTheory/Monad/Limits.lean @@ -5,7 +5,7 @@ Authors: Kim Morrison, Bhavik Mehta, Jack McKoen -/ import Mathlib.CategoryTheory.Monad.Adjunction import Mathlib.CategoryTheory.Adjunction.Limits -import Mathlib.CategoryTheory.Limits.Shapes.Terminal +import Mathlib.CategoryTheory.Limits.Shapes.IsTerminal /-! # Limits and colimits in the category of (co)algebras diff --git a/Mathlib/CategoryTheory/WithTerminal.lean b/Mathlib/CategoryTheory/WithTerminal.lean index 3ac83acd3454e..0b00c0e2f5d8f 100644 --- a/Mathlib/CategoryTheory/WithTerminal.lean +++ b/Mathlib/CategoryTheory/WithTerminal.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Adam Topaz. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Joseph Tooby-Smith, Adam Topaz -/ -import Mathlib.CategoryTheory.Limits.Shapes.Terminal +import Mathlib.CategoryTheory.Limits.Shapes.IsTerminal import Mathlib.CategoryTheory.Bicategory.Functor.Pseudofunctor /-! From 533ce927a420bd01fa3c66d250c9beb523fcd2a3 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Fri, 25 Oct 2024 15:10:12 +0000 Subject: [PATCH 420/505] feat(CategoryTheory): constructor for pseudofunctors from a strict locally discrete category (#18201) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit In this file, we define a constructor `pseudofunctorOfIsLocallyDiscrete` for the type `Pseudofunctor B C` when `C` is any bicategory, and `B` is a strict locally discrete category. Indeed, in this situation, we do not need to care about the field `map₂` of pseudofunctors because all the `2`-morphisms in `B` are identities. Co-authored-by: Joël Riou <37772949+joelriou@users.noreply.github.com> --- Mathlib.lean | 1 + .../Bicategory/Functor/LocallyDiscrete.lean | 157 ++++++++++++++++++ .../Bicategory/Grothendieck.lean | 1 + .../Bicategory/LocallyDiscrete.lean | 48 ++---- Mathlib/CategoryTheory/DiscreteCategory.lean | 20 +++ 5 files changed, 195 insertions(+), 32 deletions(-) create mode 100644 Mathlib/CategoryTheory/Bicategory/Functor/LocallyDiscrete.lean diff --git a/Mathlib.lean b/Mathlib.lean index 05b32bb4e3391..b0ddfe0730762 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1490,6 +1490,7 @@ import Mathlib.CategoryTheory.Bicategory.End import Mathlib.CategoryTheory.Bicategory.Extension import Mathlib.CategoryTheory.Bicategory.Free import Mathlib.CategoryTheory.Bicategory.Functor.Lax +import Mathlib.CategoryTheory.Bicategory.Functor.LocallyDiscrete import Mathlib.CategoryTheory.Bicategory.Functor.Oplax import Mathlib.CategoryTheory.Bicategory.Functor.Prelax import Mathlib.CategoryTheory.Bicategory.Functor.Pseudofunctor diff --git a/Mathlib/CategoryTheory/Bicategory/Functor/LocallyDiscrete.lean b/Mathlib/CategoryTheory/Bicategory/Functor/LocallyDiscrete.lean new file mode 100644 index 0000000000000..592c7e48b5551 --- /dev/null +++ b/Mathlib/CategoryTheory/Bicategory/Functor/LocallyDiscrete.lean @@ -0,0 +1,157 @@ +/- +Copyright (c) 2024 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ +import Mathlib.CategoryTheory.Bicategory.Functor.Pseudofunctor +import Mathlib.CategoryTheory.Bicategory.LocallyDiscrete + +/-! +# Pseudofunctors from locally discrete bicategories + +This file provides various ways of constructing pseudofunctors from locally discrete +bicategories. + +Firstly, we define the constructors `pseudofunctorOfIsLocallyDiscrete` and +`oplaxFunctorOfIsLocallyDiscrete` for defining pseudofunctors and oplax functors +from a locally discrete bicategories. In this situation, we do not need to care about +the field `map₂`, because all the `2`-morphisms in `B` are identities. + +We also define a specialized constructor `LocallyDiscrete.mkPseudofunctor` when +the source bicategory is of the form `B := LocallyDiscrete B₀` for a category `B₀`. + +We also prove that a functor `F : I ⥤ B` with `B` a strict bicategory can be promoted +to a pseudofunctor (or oplax functor) (`Functor.toPseudofunctor`) with domain +`LocallyDiscrete I`. + +-/ + +namespace CategoryTheory + +open Bicategory + +/-- Constructor for pseudofunctors from a locally discrete bicategory. In that +case, we do not need to provide the `map₂` field of pseudofunctors. -/ +@[simps obj map mapId mapComp] +def pseudofunctorOfIsLocallyDiscrete + {B C : Type*} [Bicategory B] [IsLocallyDiscrete B] [Bicategory C] + (obj : B → C) + (map : ∀ {b b' : B}, (b ⟶ b') → (obj b ⟶ obj b')) + (mapId : ∀ (b : B), map (𝟙 b) ≅ 𝟙 _) + (mapComp : ∀ {b₀ b₁ b₂ : B} (f : b₀ ⟶ b₁) (g : b₁ ⟶ b₂), map (f ≫ g) ≅ map f ≫ map g) + (map₂_associator : ∀ {b₀ b₁ b₂ b₃ : B} (f : b₀ ⟶ b₁) (g : b₁ ⟶ b₂) (h : b₂ ⟶ b₃), + (mapComp (f ≫ g) h).hom ≫ + (mapComp f g).hom ▷ map h ≫ (α_ (map f) (map g) (map h)).hom ≫ + map f ◁ (mapComp g h).inv ≫ (mapComp f (g ≫ h)).inv = eqToHom (by simp) := by aesop_cat) + (map₂_left_unitor : ∀ {b₀ b₁ : B} (f : b₀ ⟶ b₁), + (mapComp (𝟙 b₀) f).hom ≫ (mapId b₀).hom ▷ map f ≫ (λ_ (map f)).hom = eqToHom (by simp) := by + aesop_cat) + (map₂_right_unitor : ∀ {b₀ b₁ : B} (f : b₀ ⟶ b₁), + (mapComp f (𝟙 b₁)).hom ≫ map f ◁ (mapId b₁).hom ≫ (ρ_ (map f)).hom = eqToHom (by simp) := by + aesop_cat) : + Pseudofunctor B C where + obj := obj + map := map + map₂ φ := eqToHom (by + obtain rfl := obj_ext_of_isDiscrete φ + dsimp) + mapId := mapId + mapComp := mapComp + map₂_whisker_left _ _ _ η := by + obtain rfl := obj_ext_of_isDiscrete η + simp + map₂_whisker_right η _ := by + obtain rfl := obj_ext_of_isDiscrete η + simp + +/-- Constructor for oplax functors from a locally discrete bicategory. In that +case, we do not need to provide the `map₂` field of oplax functors. -/ +@[simps obj map mapId mapComp] +def oplaxFunctorOfIsLocallyDiscrete + {B C : Type*} [Bicategory B] [IsLocallyDiscrete B] [Bicategory C] + (obj : B → C) + (map : ∀ {b b' : B}, (b ⟶ b') → (obj b ⟶ obj b')) + (mapId : ∀ (b : B), map (𝟙 b) ⟶ 𝟙 _) + (mapComp : ∀ {b₀ b₁ b₂ : B} (f : b₀ ⟶ b₁) (g : b₁ ⟶ b₂), map (f ≫ g) ⟶ map f ≫ map g) + (map₂_associator : ∀ {b₀ b₁ b₂ b₃ : B} (f : b₀ ⟶ b₁) (g : b₁ ⟶ b₂) (h : b₂ ⟶ b₃), + eqToHom (by simp) ≫ mapComp f (g ≫ h) ≫ map f ◁ mapComp g h = + mapComp (f ≫ g) h ≫ mapComp f g ▷ map h ≫ (α_ (map f) (map g) (map h)).hom := by + aesop_cat) + (map₂_left_unitor : ∀ {b₀ b₁ : B} (f : b₀ ⟶ b₁), + mapComp (𝟙 b₀) f ≫ mapId b₀ ▷ map f ≫ (λ_ (map f)).hom = eqToHom (by simp) := by + aesop_cat) + (map₂_right_unitor : ∀ {b₀ b₁ : B} (f : b₀ ⟶ b₁), + mapComp f (𝟙 b₁) ≫ map f ◁ mapId b₁ ≫ (ρ_ (map f)).hom = eqToHom (by simp) := by + aesop_cat) : + OplaxFunctor B C where + obj := obj + map := map + map₂ φ := eqToHom (by + obtain rfl := obj_ext_of_isDiscrete φ + dsimp) + mapId := mapId + mapComp := mapComp + mapComp_naturality_left η := by + obtain rfl := obj_ext_of_isDiscrete η + simp + mapComp_naturality_right _ _ _ η := by + obtain rfl := obj_ext_of_isDiscrete η + simp + +section + +variable {I B : Type*} [Category I] [Bicategory B] [Strict B] (F : I ⥤ B) + +attribute [local simp] + Strict.leftUnitor_eqToIso Strict.rightUnitor_eqToIso Strict.associator_eqToIso + +/-- +If `B` is a strict bicategory and `I` is a (1-)category, any functor (of 1-categories) `I ⥤ B` can +be promoted to a pseudofunctor from `LocallyDiscrete I` to `B`. +-/ +@[simps! obj map mapId mapComp] +def Functor.toPseudoFunctor : Pseudofunctor (LocallyDiscrete I) B := + pseudofunctorOfIsLocallyDiscrete + (fun ⟨X⟩ ↦ F.obj X) (fun ⟨f⟩ ↦ F.map f) + (fun ⟨X⟩ ↦ eqToIso (by simp)) (fun f g ↦ eqToIso (by simp)) + +/-- +If `B` is a strict bicategory and `I` is a (1-)category, any functor (of 1-categories) `I ⥤ B` can +be promoted to an oplax functor from `LocallyDiscrete I` to `B`. +-/ +@[simps! obj map mapId mapComp] +def Functor.toOplaxFunctor : OplaxFunctor (LocallyDiscrete I) B := + oplaxFunctorOfIsLocallyDiscrete + (fun ⟨X⟩ ↦ F.obj X) (fun ⟨f⟩ ↦ F.map f) + (fun ⟨X⟩ ↦ eqToHom (by simp)) (fun f g ↦ eqToHom (by simp)) + +end + +namespace LocallyDiscrete + +/-- Constructor for pseudofunctors from a locally discrete bicategory. In that +case, we do not need to provide the `map₂` field of pseudofunctors. -/ +@[simps! obj map mapId mapComp] +def mkPseudofunctor {B₀ C : Type*} [Category B₀] [Bicategory C] + (obj : B₀ → C) + (map : ∀ {b b' : B₀}, (b ⟶ b') → (obj b ⟶ obj b')) + (mapId : ∀ (b : B₀), map (𝟙 b) ≅ 𝟙 _) + (mapComp : ∀ {b₀ b₁ b₂ : B₀} (f : b₀ ⟶ b₁) (g : b₁ ⟶ b₂), map (f ≫ g) ≅ map f ≫ map g) + (map₂_associator : ∀ {b₀ b₁ b₂ b₃ : B₀} (f : b₀ ⟶ b₁) (g : b₁ ⟶ b₂) (h : b₂ ⟶ b₃), + (mapComp (f ≫ g) h).hom ≫ + (mapComp f g).hom ▷ map h ≫ (α_ (map f) (map g) (map h)).hom ≫ + map f ◁ (mapComp g h).inv ≫ (mapComp f (g ≫ h)).inv = eqToHom (by simp) := by aesop_cat) + (map₂_left_unitor : ∀ {b₀ b₁ : B₀} (f : b₀ ⟶ b₁), + (mapComp (𝟙 b₀) f).hom ≫ (mapId b₀).hom ▷ map f ≫ (λ_ (map f)).hom = eqToHom (by simp) := by + aesop_cat) + (map₂_right_unitor : ∀ {b₀ b₁ : B₀} (f : b₀ ⟶ b₁), + (mapComp f (𝟙 b₁)).hom ≫ map f ◁ (mapId b₁).hom ≫ (ρ_ (map f)).hom = eqToHom (by simp) := by + aesop_cat) : + Pseudofunctor (LocallyDiscrete B₀) C := + pseudofunctorOfIsLocallyDiscrete (fun b ↦ obj b.as) (fun f ↦ map f.as) + (fun _ ↦ mapId _) (fun _ _ ↦ mapComp _ _) (fun _ _ _ ↦ map₂_associator _ _ _) + (fun _ ↦ map₂_left_unitor _) (fun _ ↦ map₂_right_unitor _) + +end LocallyDiscrete + +end CategoryTheory diff --git a/Mathlib/CategoryTheory/Bicategory/Grothendieck.lean b/Mathlib/CategoryTheory/Bicategory/Grothendieck.lean index d7eb112f1f6ba..efb94ebc6c24a 100644 --- a/Mathlib/CategoryTheory/Bicategory/Grothendieck.lean +++ b/Mathlib/CategoryTheory/Bicategory/Grothendieck.lean @@ -5,6 +5,7 @@ Authors: Calle Sönne -/ import Mathlib.CategoryTheory.Bicategory.LocallyDiscrete +import Mathlib.CategoryTheory.Bicategory.Functor.Pseudofunctor /-! # The Grothendieck construction diff --git a/Mathlib/CategoryTheory/Bicategory/LocallyDiscrete.lean b/Mathlib/CategoryTheory/Bicategory/LocallyDiscrete.lean index 2e93d740e778d..7db3d423d41ad 100644 --- a/Mathlib/CategoryTheory/Bicategory/LocallyDiscrete.lean +++ b/Mathlib/CategoryTheory/Bicategory/LocallyDiscrete.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yuma Mizuno, Calle Sönne -/ import Mathlib.CategoryTheory.DiscreteCategory -import Mathlib.CategoryTheory.Bicategory.Functor.Pseudofunctor +import Mathlib.CategoryTheory.Bicategory.Functor.Prelax import Mathlib.CategoryTheory.Bicategory.Strict /-! @@ -21,8 +21,6 @@ namespace CategoryTheory open Bicategory Discrete -open Bicategory - universe w₂ w₁ v₂ v₁ v u₂ u₁ u section @@ -106,35 +104,6 @@ instance locallyDiscreteBicategory.strict : Strict (LocallyDiscrete C) where comp_id _ := Discrete.ext (Category.comp_id _) assoc _ _ _ := Discrete.ext (Category.assoc _ _ _) -attribute [local simp] - Strict.leftUnitor_eqToIso Strict.rightUnitor_eqToIso Strict.associator_eqToIso - -variable {I : Type u₁} [Category.{v₁} I] {B : Type u₂} [Bicategory.{w₂, v₂} B] [Strict B] - -/-- -If `B` is a strict bicategory and `I` is a (1-)category, any functor (of 1-categories) `I ⥤ B` can -be promoted to a pseudofunctor from `LocallyDiscrete I` to `B`. --/ -@[simps] -def Functor.toPseudoFunctor (F : I ⥤ B) : Pseudofunctor (LocallyDiscrete I) B where - obj i := F.obj i.as - map f := F.map f.as - map₂ η := eqToHom (congr_arg _ (LocallyDiscrete.eq_of_hom η)) - mapId i := eqToIso (F.map_id i.as) - mapComp f g := eqToIso (F.map_comp f.as g.as) - -/-- -If `B` is a strict bicategory and `I` is a (1-)category, any functor (of 1-categories) `I ⥤ B` can -be promoted to an oplax functor from `LocallyDiscrete I` to `B`. --/ -@[simps] -def Functor.toOplaxFunctor (F : I ⥤ B) : OplaxFunctor (LocallyDiscrete I) B where - obj i := F.obj i.as - map f := F.map f.as - map₂ η := eqToHom (congr_arg _ (LocallyDiscrete.eq_of_hom η)) - mapId i := eqToHom (F.map_id i.as) - mapComp f g := eqToHom (F.map_comp f.as g.as) - end section @@ -148,6 +117,21 @@ lemma PrelaxFunctor.map₂_eqToHom (F : PrelaxFunctor B C) {a b : B} {f g : a end +namespace Bicategory + +/-- A bicategory is locally discrete if the categories of 1-morphisms are discrete. -/ +abbrev IsLocallyDiscrete (B : Type*) [Bicategory B] := ∀ (b c : B), IsDiscrete (b ⟶ c) + +instance (C : Type*) [Category C] : IsLocallyDiscrete (LocallyDiscrete C) := + fun _ _ ↦ Discrete.isDiscrete _ + +instance (B : Type*) [Bicategory B] [IsLocallyDiscrete B] : Strict B where + id_comp f := obj_ext_of_isDiscrete (leftUnitor f).hom + comp_id f := obj_ext_of_isDiscrete (rightUnitor f).hom + assoc f g h := obj_ext_of_isDiscrete (associator f g h).hom + +end Bicategory + end CategoryTheory section diff --git a/Mathlib/CategoryTheory/DiscreteCategory.lean b/Mathlib/CategoryTheory/DiscreteCategory.lean index a6f0f73828116..1d28f0410aa9e 100644 --- a/Mathlib/CategoryTheory/DiscreteCategory.lean +++ b/Mathlib/CategoryTheory/DiscreteCategory.lean @@ -304,4 +304,24 @@ def piEquivalenceFunctorDiscrete (J : Type u₂) (C : Type u₁) [Category.{v₁ obtain rfl : f = 𝟙 _ := rfl simp))) (by aesop_cat) +/-- A category is discrete when there is at most one morphism between two objects, +in which case they are equal. -/ +class IsDiscrete (C : Type*) [Category C] : Prop where + subsingleton (X Y : C) : Subsingleton (X ⟶ Y) := by infer_instance + eq_of_hom {X Y : C} (f : X ⟶ Y) : X = Y + +attribute [instance] IsDiscrete.subsingleton + +lemma obj_ext_of_isDiscrete {C : Type*} [Category C] [IsDiscrete C] + {X Y : C} (f : X ⟶ Y) : X = Y := IsDiscrete.eq_of_hom f + +instance Discrete.isDiscrete (C : Type*) : IsDiscrete (Discrete C) where + eq_of_hom := by rintro ⟨_⟩ ⟨_⟩ ⟨⟨rfl⟩⟩; rfl + +instance (C : Type*) [Category C] [IsDiscrete C] : IsDiscrete Cᵒᵖ where + eq_of_hom := by + rintro ⟨_⟩ ⟨_⟩ ⟨f⟩ + obtain rfl := obj_ext_of_isDiscrete f + rfl + end CategoryTheory From 7f0670510ccf43e0e595357e31b7b8212cc02c45 Mon Sep 17 00:00:00 2001 From: Riccardo Brasca Date: Fri, 25 Oct 2024 15:28:41 +0000 Subject: [PATCH 421/505] feat: add pow_sub_one_dvd_differentIdeal (#18228) From flt-regular. --- .../RingTheory/DedekindDomain/Different.lean | 69 +++++++++++++++++++ Mathlib/RingTheory/Trace/Basic.lean | 19 +++++ 2 files changed, 88 insertions(+) diff --git a/Mathlib/RingTheory/DedekindDomain/Different.lean b/Mathlib/RingTheory/DedekindDomain/Different.lean index 915f686f016f1..a9e9ad8924c5b 100644 --- a/Mathlib/RingTheory/DedekindDomain/Different.lean +++ b/Mathlib/RingTheory/DedekindDomain/Different.lean @@ -7,6 +7,8 @@ import Mathlib.RingTheory.DedekindDomain.Ideal import Mathlib.RingTheory.Discriminant import Mathlib.RingTheory.DedekindDomain.IntegralClosure import Mathlib.NumberTheory.KummerDedekind +import Mathlib.RingTheory.IntegralClosure.IntegralRestrict +import Mathlib.RingTheory.Trace.Quotient /-! # The different ideal @@ -568,3 +570,70 @@ lemma aeval_derivative_mem_differentIdeal [NoZeroSMulDivisors A B] refine SetLike.le_def.mp ?_ (Ideal.mem_span_singleton_self _) rw [← conductor_mul_differentIdeal A K L x hx] exact Ideal.mul_le_left + +section + +include K L in +lemma pow_sub_one_dvd_differentIdeal_aux [IsFractionRing B L] [IsDedekindDomain A] + [NoZeroSMulDivisors A B] [Module.Finite A B] + {p : Ideal A} [p.IsMaximal] (P : Ideal B) {e : ℕ} (he : e ≠ 0) (hp : p ≠ ⊥) + (hP : P ^ e ∣ p.map (algebraMap A B)) : P ^ (e - 1) ∣ differentIdeal A B := by + obtain ⟨a, ha⟩ := (pow_dvd_pow _ (Nat.sub_le e 1)).trans hP + have hp' := (Ideal.map_eq_bot_iff_of_injective + (NoZeroSMulDivisors.algebraMap_injective A B)).not.mpr hp + have habot : a ≠ ⊥ := fun ha' ↦ hp' (by simpa [ha'] using ha) + have hPbot : P ≠ ⊥ := by + rintro rfl; apply hp' + rwa [← Ideal.zero_eq_bot, zero_pow he, zero_dvd_iff, Ideal.zero_eq_bot] at hP + have : p.map (algebraMap A B) ∣ a ^ e := by + obtain ⟨b, hb⟩ := hP + apply_fun (· ^ e : Ideal B → _) at ha + apply_fun (· ^ (e - 1) : Ideal B → _) at hb + simp only [mul_pow, ← pow_mul, mul_comm e] at ha hb + conv_lhs at ha => rw [← Nat.sub_add_cancel (Nat.one_le_iff_ne_zero.mpr he)] + rw [pow_add, hb, mul_assoc, mul_right_inj' (pow_ne_zero _ hPbot), pow_one, mul_comm] at ha + exact ⟨_, ha.symm⟩ + suffices ∀ x ∈ a, intTrace A B x ∈ p by + have hP : ((P ^ (e - 1) : _)⁻¹ : FractionalIdeal B⁰ L) = a / p.map (algebraMap A B) := by + apply inv_involutive.injective + simp only [inv_inv, ha, FractionalIdeal.coeIdeal_mul, inv_div, ne_eq, + FractionalIdeal.coeIdeal_eq_zero, mul_div_assoc] + rw [div_self (by simpa), mul_one] + rw [Ideal.dvd_iff_le, differentialIdeal_le_iff (K := K) (L := L) (pow_ne_zero _ hPbot), hP, + Submodule.map_le_iff_le_comap] + intro x hx + rw [Submodule.restrictScalars_mem, FractionalIdeal.mem_coe, + FractionalIdeal.mem_div_iff_of_nonzero (by simpa using hp')] at hx + rw [Submodule.mem_comap, LinearMap.coe_restrictScalars, ← FractionalIdeal.coe_one, + ← div_self (G₀ := FractionalIdeal A⁰ K) (a := p) (by simpa using hp), + FractionalIdeal.mem_coe, FractionalIdeal.mem_div_iff_of_nonzero (by simpa using hp)] + simp only [FractionalIdeal.mem_coeIdeal, forall_exists_index, and_imp, + forall_apply_eq_imp_iff₂] at hx + intro y hy' + obtain ⟨y, hy, rfl : algebraMap A K _ = _⟩ := (FractionalIdeal.mem_coeIdeal _).mp hy' + obtain ⟨z, hz, hz'⟩ := hx _ (Ideal.mem_map_of_mem _ hy) + have : trace K L (algebraMap B L z) ∈ (p : FractionalIdeal A⁰ K) := by + rw [← algebraMap_intTrace (A := A)] + exact ⟨intTrace A B z, this z hz, rfl⟩ + rwa [mul_comm, ← smul_eq_mul, ← LinearMap.map_smul, smul_def, mul_comm, + ← IsScalarTower.algebraMap_apply, IsScalarTower.algebraMap_apply A B L, ← hz'] + intros x hx + rw [← Ideal.Quotient.eq_zero_iff_mem, ← trace_quotient_eq_of_isDedekindDomain, + ← isNilpotent_iff_eq_zero] + refine trace_isNilpotent_of_isNilpotent ⟨e, ?_⟩ + rw [← map_pow, Ideal.Quotient.eq_zero_iff_mem] + exact (Ideal.dvd_iff_le.mp this) <| Ideal.pow_mem_pow hx _ + +lemma pow_sub_one_dvd_differentIdeal [IsDedekindDomain A] [NoZeroSMulDivisors A B] + [Module.Finite A B] [Algebra.IsSeparable (FractionRing A) (FractionRing B)] + {p : Ideal A} [p.IsMaximal] (P : Ideal B) (e : ℕ) (hp : p ≠ ⊥) + (hP : P ^ e ∣ p.map (algebraMap A B)) : P ^ (e - 1) ∣ differentIdeal A B := by + have : IsLocalization (algebraMapSubmonoid B A⁰) (FractionRing B) := + IsIntegralClosure.isLocalization _ (FractionRing A) _ _ + have : FiniteDimensional (FractionRing A) (FractionRing B) := + Module.Finite_of_isLocalization A B _ _ A⁰ + by_cases he : e = 0 + · rw [he, pow_zero]; exact one_dvd _ + exact pow_sub_one_dvd_differentIdeal_aux A (FractionRing A) (FractionRing B) _ he hp hP + +end diff --git a/Mathlib/RingTheory/Trace/Basic.lean b/Mathlib/RingTheory/Trace/Basic.lean index e06b0141ff923..5d15d4fac0ac3 100644 --- a/Mathlib/RingTheory/Trace/Basic.lean +++ b/Mathlib/RingTheory/Trace/Basic.lean @@ -483,3 +483,22 @@ lemma traceForm_dualBasis_powerBasis_eq [FiniteDimensional K L] [Algebra.IsSepar ring end DetNeZero + +section isNilpotent + +namespace Algebra + +/-- The trace of a nilpotent element is nilpotent. -/ +lemma trace_isNilpotent_of_isNilpotent {R S : Type*} [CommRing R] [CommRing S] [Algebra R S] {x : S} + (hx : IsNilpotent x) : IsNilpotent (trace R S x) := by + by_cases hS : ∃ s : Finset S, Nonempty (Basis s R S) + · obtain ⟨s, ⟨b⟩⟩ := hS + have := Module.Finite.of_basis b + have := (Module.free_def R S).mpr ⟨s, ⟨b⟩⟩ + apply LinearMap.isNilpotent_trace_of_isNilpotent (hx.map (lmul R S)) + · rw [trace_eq_zero_of_not_exists_basis _ hS, LinearMap.zero_apply] + exact IsNilpotent.zero + +end Algebra + +end isNilpotent From 159c3b56230633e79ccee2399eab15ebe2284a0e Mon Sep 17 00:00:00 2001 From: Scott Carnahan <128885296+ScottCarnahan@users.noreply.github.com> Date: Fri, 25 Oct 2024 17:41:50 +0000 Subject: [PATCH 422/505] feat (RingTheory/HahnSeries/Summable) : Pointwise SMul for summable families (#16701) This PR adds a definition and some API for pointwise SMul between summable families. This generalizes the action of a single Hahn series on a summable family. - [ ] depends on: #16649 [sundry API] - [ ] depends on: #16871 [single summable family] - [ ] depends on: #16872 [Equiv of families] --- Mathlib/RingTheory/HahnSeries/Summable.lean | 215 ++++++++++++++------ 1 file changed, 148 insertions(+), 67 deletions(-) diff --git a/Mathlib/RingTheory/HahnSeries/Summable.lean b/Mathlib/RingTheory/HahnSeries/Summable.lean index c3b6bf05ac681..c9ea82db6087b 100644 --- a/Mathlib/RingTheory/HahnSeries/Summable.lean +++ b/Mathlib/RingTheory/HahnSeries/Summable.lean @@ -3,7 +3,6 @@ Copyright (c) 2021 Aaron Anderson. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Aaron Anderson -/ -import Mathlib.Algebra.BigOperators.Finprod import Mathlib.RingTheory.HahnSeries.Multiplication /-! @@ -29,7 +28,6 @@ commutative domain. * Remove unnecessary domain hypotheses. * More general summable families, e.g., define the evaluation homomorphism from a power series ring taking `X` to a positive order element. - * Generalize `SMul` to Hahn modules. ## References - [J. van der Hoeven, *Operators on Generalized Power Series*][van_der_hoeven] @@ -42,7 +40,7 @@ open Pointwise noncomputable section -variable {Γ R α β : Type*} +variable {Γ Γ' R V α β : Type*} namespace HahnSeries @@ -232,7 +230,7 @@ end AddCommMonoid section AddCommGroup -variable [PartialOrder Γ] [AddCommGroup R] {α : Type*} {s t : SummableFamily Γ R α} {a : α} +variable [PartialOrder Γ] [AddCommGroup R] {s t : SummableFamily Γ R α} {a : α} instance : Neg (SummableFamily Γ R α) := ⟨fun s => @@ -267,73 +265,155 @@ theorem sub_apply : (s - t) a = s a - t a := end AddCommGroup -section Semiring - -variable [OrderedCancelAddCommMonoid Γ] [Semiring R] {α : Type*} - -instance : SMul (HahnSeries Γ R) (SummableFamily Γ R α) where - smul x s := - { toFun := fun a => x * s a - isPWO_iUnion_support' := by - apply (x.isPWO_support.add s.isPWO_iUnion_support).mono - refine Set.Subset.trans (Set.iUnion_mono fun a => support_mul_subset_add_support) ?_ - intro g - simp only [Set.mem_iUnion, exists_imp] - exact fun a ha => (Set.add_subset_add (Set.Subset.refl _) (Set.subset_iUnion _ a)) ha - finite_co_support' := fun g => by - apply ((addAntidiagonal x.isPWO_support s.isPWO_iUnion_support g).finite_toSet.biUnion' - fun ij _ => ?_).subset fun a ha => ?_ - · exact fun ij _ => Function.support fun a => (s a).coeff ij.2 - · apply s.finite_co_support - · obtain ⟨i, hi, j, hj, rfl⟩ := support_mul_subset_add_support ha - simp only [exists_prop, Set.mem_iUnion, mem_addAntidiagonal, mul_coeff, mem_support, - isPWO_support, Prod.exists] - exact ⟨i, j, mem_coe.2 (mem_addAntidiagonal.2 ⟨hi, Set.mem_iUnion.2 ⟨a, hj⟩, rfl⟩), hj⟩ } +section SMul + +variable [PartialOrder Γ] [PartialOrder Γ'] [AddCommMonoid R] [AddCommMonoid V] [SMulWithZero R V] + +theorem smul_support_subset_prod (s : SummableFamily Γ R α) + (t : SummableFamily Γ' V β) (gh : Γ × Γ') : + (Function.support fun (i : α × β) ↦ (s i.1).coeff gh.1 • (t i.2).coeff gh.2) ⊆ + ((s.finite_co_support' gh.1).prod (t.finite_co_support' gh.2)).toFinset := by + intro _ hab + simp_all only [Function.mem_support, ne_eq, Set.Finite.coe_toFinset, Set.mem_prod, + Set.mem_setOf_eq] + exact ⟨left_ne_zero_of_smul hab, right_ne_zero_of_smul hab⟩ + +theorem smul_support_finite (s : SummableFamily Γ R α) + (t : SummableFamily Γ' V β) (gh : Γ × Γ') : + (Function.support fun (i : α × β) ↦ (s i.1).coeff gh.1 • (t i.2).coeff gh.2).Finite := + Set.Finite.subset (Set.toFinite ((s.finite_co_support' gh.1).prod + (t.finite_co_support' gh.2)).toFinset) (smul_support_subset_prod s t gh) + +variable [VAdd Γ Γ'] [IsOrderedCancelVAdd Γ Γ'] + +open HahnModule + +theorem isPWO_iUnion_support_prod_smul {s : α → HahnSeries Γ R} {t : β → HahnSeries Γ' V} + (hs : (⋃ a, (s a).support).IsPWO) (ht : (⋃ b, (t b).support).IsPWO) : + (⋃ (a : α × β), ((fun a ↦ (of R).symm + ((s a.1) • (of R) (t a.2))) a).support).IsPWO := by + apply (hs.vadd ht).mono + have hsupp : ∀ ab : α × β, support ((fun ab ↦ (of R).symm (s ab.1 • (of R) (t ab.2))) ab) ⊆ + (s ab.1).support +ᵥ (t ab.2).support := by + intro ab + refine Set.Subset.trans (fun x hx => ?_) (support_vaddAntidiagonal_subset_vadd + (hs := (s ab.1).isPWO_support) (ht := (t ab.2).isPWO_support)) + contrapose! hx + simp only [Set.mem_setOf_eq, not_nonempty_iff_eq_empty] at hx + rw [mem_support, not_not, HahnModule.smul_coeff, hx, sum_empty] + refine Set.Subset.trans (Set.iUnion_mono fun a => (hsupp a)) ?_ + simp_all only [Set.iUnion_subset_iff, Prod.forall] + exact fun a b => Set.vadd_subset_vadd (Set.subset_iUnion_of_subset a fun x y ↦ y) + (Set.subset_iUnion_of_subset b fun x y ↦ y) + +theorem finite_co_support_prod_smul (s : SummableFamily Γ R α) + (t : SummableFamily Γ' V β) (g : Γ') : + Finite {(ab : α × β) | + ((fun (ab : α × β) ↦ (of R).symm (s ab.1 • (of R) (t ab.2))) ab).coeff g ≠ 0} := by + apply ((VAddAntidiagonal s.isPWO_iUnion_support t.isPWO_iUnion_support g).finite_toSet.biUnion' + (fun gh _ => smul_support_finite s t gh)).subset _ + exact fun ab hab => by + simp only [smul_coeff, ne_eq, Set.mem_setOf_eq] at hab + obtain ⟨ij, hij⟩ := Finset.exists_ne_zero_of_sum_ne_zero hab + simp only [mem_coe, mem_vaddAntidiagonal, Set.mem_iUnion, mem_support, ne_eq, + Function.mem_support, exists_prop, Prod.exists] + exact ⟨ij.1, ij.2, ⟨⟨ab.1, left_ne_zero_of_smul hij.2⟩, ⟨ab.2, right_ne_zero_of_smul hij.2⟩, + ((mem_vaddAntidiagonal _ _ _).mp hij.1).2.2⟩, hij.2⟩ + +/-- An elementwise scalar multiplication of one summable family on another. -/ +@[simps] +def FamilySMul (s : SummableFamily Γ R α) (t : SummableFamily Γ' V β) : + (SummableFamily Γ' V (α × β)) where + toFun ab := (of R).symm (s (ab.1) • ((of R) (t (ab.2)))) + isPWO_iUnion_support' := + isPWO_iUnion_support_prod_smul s.isPWO_iUnion_support t.isPWO_iUnion_support + finite_co_support' g := finite_co_support_prod_smul s t g + +theorem sum_vAddAntidiagonal_eq (s : SummableFamily Γ R α) (t : SummableFamily Γ' V β) (g : Γ') + (a : α × β) : + ∑ x ∈ VAddAntidiagonal (s a.1).isPWO_support' (t a.2).isPWO_support' g, (s a.1).coeff x.1 • + (t a.2).coeff x.2 = ∑ x ∈ VAddAntidiagonal s.isPWO_iUnion_support' t.isPWO_iUnion_support' g, + (s a.1).coeff x.1 • (t a.2).coeff x.2 := by + refine sum_subset (fun gh hgh => ?_) fun gh hgh h => ?_ + · simp_all only [mem_vaddAntidiagonal, Function.mem_support, Set.mem_iUnion, mem_support] + exact ⟨Exists.intro a.1 hgh.1, Exists.intro a.2 hgh.2.1, trivial⟩ + · by_cases hs : (s a.1).coeff gh.1 = 0 + · exact smul_eq_zero_of_left hs ((t a.2).coeff gh.2) + · simp_all + +theorem family_smul_coeff {R} {V} [Semiring R] [AddCommMonoid V] [Module R V] + (s : SummableFamily Γ R α) (t : SummableFamily Γ' V β) (g : Γ') : + (FamilySMul s t).hsum.coeff g = ∑ gh ∈ VAddAntidiagonal s.isPWO_iUnion_support + t.isPWO_iUnion_support g, (s.hsum.coeff gh.1) • (t.hsum.coeff gh.2) := by + rw [hsum_coeff] + simp only [hsum_coeff_eq_sum, FamilySMul_toFun, HahnModule.smul_coeff, Equiv.symm_apply_apply] + simp_rw [sum_vAddAntidiagonal_eq, Finset.smul_sum, Finset.sum_smul] + rw [← sum_finsum_comm _ _ <| fun gh _ => smul_support_finite s t gh] + refine sum_congr rfl fun gh _ => ?_ + rw [finsum_eq_sum _ (smul_support_finite s t gh), ← sum_product_right'] + refine sum_subset (fun ab hab => ?_) (fun ab _ hab => by simp_all) + have hsupp := smul_support_subset_prod s t gh + simp_all only [mem_vaddAntidiagonal, Set.mem_iUnion, mem_support, ne_eq, Set.Finite.mem_toFinset, + Function.mem_support, Set.Finite.coe_toFinset, support_subset_iff, Set.mem_prod, + Set.mem_setOf_eq, Prod.forall, coeff_support, mem_product] + exact hsupp ab.1 ab.2 hab + +theorem hsum_family_smul {R} {V} [Semiring R] [AddCommMonoid V] [Module R V] + (s : SummableFamily Γ R α) (t : SummableFamily Γ' V β) : + (FamilySMul s t).hsum = (of R).symm (s.hsum • (of R) (t.hsum)) := by + ext g + rw [family_smul_coeff s t g, HahnModule.smul_coeff, Equiv.symm_apply_apply] + refine Eq.symm (sum_of_injOn (fun a ↦ a) (fun _ _ _ _ h ↦ h) (fun _ hgh => ?_) + (fun gh _ hgh => ?_) fun _ _ => by simp) + · simp_all only [mem_coe, mem_vaddAntidiagonal, mem_support, ne_eq, Set.mem_iUnion, and_true] + constructor + · rw [hsum_coeff_eq_sum] at hgh + have h' := Finset.exists_ne_zero_of_sum_ne_zero hgh.1 + simpa using h' + · by_contra hi + simp_all + · simp only [Set.image_id', mem_coe, mem_vaddAntidiagonal, mem_support, ne_eq, not_and] at hgh + by_cases h : s.hsum.coeff gh.1 = 0 + · exact smul_eq_zero_of_left h (t.hsum.coeff gh.2) + · simp_all + +instance [AddCommMonoid R] [SMulWithZero R V] : SMul (HahnSeries Γ R) (SummableFamily Γ' V β) where + smul x t := Equiv (Equiv.punitProd β) <| FamilySMul (single x) t + +theorem smul_eq {x : HahnSeries Γ R} {t : SummableFamily Γ' V β} : + x • t = Equiv (Equiv.punitProd β) (FamilySMul (single x) t) := + rfl @[simp] -theorem smul_apply {x : HahnSeries Γ R} {s : SummableFamily Γ R α} {a : α} : (x • s) a = x * s a := +theorem smul_apply {x : HahnSeries Γ R} {s : SummableFamily Γ' V α} {a : α} : + (x • s) a = (of R).symm (x • of R (s a)) := rfl -instance : Module (HahnSeries Γ R) (SummableFamily Γ R α) where +@[simp] +theorem hsum_smul_module {R} {V} [Semiring R] [AddCommMonoid V] [Module R V] {x : HahnSeries Γ R} + {s : SummableFamily Γ' V α} : + (x • s).hsum = (of R).symm (x • of R s.hsum) := by + rw [smul_eq, hsum_equiv, hsum_family_smul, hsum_single] + +end SMul + +section Semiring + +variable [OrderedCancelAddCommMonoid Γ] [PartialOrder Γ'] [AddAction Γ Γ'] + [IsOrderedCancelVAdd Γ Γ'] [Semiring R] + +instance [AddCommMonoid V] [Module R V] : Module (HahnSeries Γ R) (SummableFamily Γ' V α) where smul := (· • ·) - smul_zero _ := ext fun _ => mul_zero _ - zero_smul _ := ext fun _ => zero_mul _ - one_smul _ := ext fun _ => one_mul _ - add_smul _ _ _ := ext fun _ => add_mul _ _ _ - smul_add _ _ _ := ext fun _ => mul_add _ _ _ - mul_smul _ _ _ := ext fun _ => mul_assoc _ _ _ + smul_zero _ := ext fun _ => by simp + zero_smul _ := ext fun _ => by simp + one_smul _ := ext fun _ => by rw [smul_apply, HahnModule.one_smul', Equiv.symm_apply_apply] + add_smul _ _ _ := ext fun _ => by simp [add_smul] + smul_add _ _ _ := ext fun _ => by simp + mul_smul _ _ _ := ext fun _ => by simp [HahnModule.instModule.mul_smul] -@[simp] -theorem hsum_smul {x : HahnSeries Γ R} {s : SummableFamily Γ R α} : (x • s).hsum = x * s.hsum := by - ext g - simp only [mul_coeff, hsum_coeff, smul_apply] - refine - (Eq.trans (finsum_congr fun a => ?_) - (finsum_sum_comm (addAntidiagonal x.isPWO_support s.isPWO_iUnion_support g) - (fun i ij => x.coeff (Prod.fst ij) * (s i).coeff ij.snd) ?_)).trans - ?_ - · refine sum_subset (addAntidiagonal_mono_right - (Set.subset_iUnion (fun j => support (toFun s j)) a)) ?_ - rintro ⟨i, j⟩ hU ha - rw [mem_addAntidiagonal] at * - rw [Classical.not_not.1 fun con => ha ⟨hU.1, con, hU.2.2⟩, mul_zero] - · rintro ⟨i, j⟩ _ - refine (s.finite_co_support j).subset ?_ - simp_rw [Function.support_subset_iff', Function.mem_support, Classical.not_not] - intro a ha - rw [ha, mul_zero] - · refine (sum_congr rfl ?_).trans (sum_subset (addAntidiagonal_mono_right ?_) ?_).symm - · rintro ⟨i, j⟩ _ - rw [mul_finsum] - apply s.finite_co_support - · intro x hx - simp only [Set.mem_iUnion, Ne, mem_support] - contrapose! hx - simp [hx] - · rintro ⟨i, j⟩ hU ha - rw [mem_addAntidiagonal] at * - rw [← hsum_coeff, Classical.not_not.1 fun con => ha ⟨hU.1, con, hU.2.2⟩, - mul_zero] +theorem hsum_smul {x : HahnSeries Γ R} {s : SummableFamily Γ R α} : + (x • s).hsum = x * s.hsum := by + rw [hsum_smul_module, of_symm_smul_of_eq_mul] /-- The summation of a `summable_family` as a `LinearMap`. -/ @[simps] @@ -441,6 +521,7 @@ section powers variable [LinearOrderedCancelAddCommMonoid Γ] [CommRing R] [IsDomain R] /-- The powers of an element of positive valuation form a summable family. -/ +@[simps] def powers (x : HahnSeries Γ R) (hx : 0 < x.orderTop) : SummableFamily Γ R ℕ where toFun n := x ^ n isPWO_iUnion_support' := isPWO_iUnion_support_powers hx @@ -486,8 +567,8 @@ theorem embDomain_succ_smul_powers : rw [Set.mem_range, not_exists] exact Nat.succ_ne_zero · refine Eq.trans (embDomain_image _ ⟨Nat.succ, Nat.succ_injective⟩) ?_ - simp only [pow_succ', coe_powers, coe_sub, smul_apply, coe_ofFinsupp, Pi.sub_apply] - rw [Finsupp.single_eq_of_ne n.succ_ne_zero.symm, sub_zero] + rw [smul_apply, powers_toFun, coe_sub, coe_powers, Pi.sub_apply, coe_ofFinsupp, pow_succ', + Finsupp.single_eq_of_ne (Nat.zero_ne_add_one n), sub_zero, of_symm_smul_of_eq_mul] theorem one_sub_self_mul_hsum_powers : (1 - x) * (powers x hx).hsum = 1 := by rw [← hsum_smul, sub_smul 1 x (powers x hx), one_smul, hsum_sub, ← From 6ac1e12b7a0ab36e6738771fb6dfaa8bac131916 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Fri, 25 Oct 2024 18:25:24 +0000 Subject: [PATCH 423/505] feat: add `isCompact_{quasi}spectrum` for instances of the continuous functional calculus (#18117) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We already have instances for `CompactSpace`, but it's handy to have these when you need to work with the sets themselves (e.g., to show that `spectrum R a ∪ spectrum R b` is compact). --- .../CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean | 5 +++++ .../CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean | 5 +++++ 2 files changed, 10 insertions(+) diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean index 73105d7530e9b..ed76c4a88dc63 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean @@ -106,6 +106,11 @@ variable [TopologicalSemiring R] [ContinuousStar R] [NonUnitalRing A] [StarRing variable [TopologicalSpace A] [Module R A] [IsScalarTower R A A] [SMulCommClass R A A] variable [instCFCₙ : NonUnitalContinuousFunctionalCalculus R p] +include instCFCₙ in +lemma NonUnitalContinuousFunctionalCalculus.isCompact_quasispectrum (a : A) : + IsCompact (σₙ R a) := + isCompact_iff_compactSpace.mpr inferInstance + lemma NonUnitalStarAlgHom.ext_continuousMap [UniqueNonUnitalContinuousFunctionalCalculus R A] (a : A) (φ ψ : C(σₙ R a, R)₀ →⋆ₙₐ[R] A) (hφ : Continuous φ) (hψ : Continuous ψ) (h : φ ⟨.restrict (σₙ R a) <| .id R, rfl⟩ = ψ ⟨.restrict (σₙ R a) <| .id R, rfl⟩) : diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean index b63d226c61d6d..b193f3c40b439 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean @@ -200,6 +200,11 @@ variable {R A : Type*} {p : A → Prop} [CommSemiring R] [StarRing R] [MetricSpa variable [TopologicalSemiring R] [ContinuousStar R] [TopologicalSpace A] [Ring A] [StarRing A] variable [Algebra R A] [instCFC : ContinuousFunctionalCalculus R p] +include instCFC in +lemma ContinuousFunctionalCalculus.isCompact_spectrum (a : A) : + IsCompact (spectrum R a) := + isCompact_iff_compactSpace.mpr inferInstance + lemma StarAlgHom.ext_continuousMap [UniqueContinuousFunctionalCalculus R A] (a : A) (φ ψ : C(spectrum R a, R) →⋆ₐ[R] A) (hφ : Continuous φ) (hψ : Continuous ψ) (h : φ (.restrict (spectrum R a) <| .id R) = ψ (.restrict (spectrum R a) <| .id R)) : From 970d241d0868926adcfcc688a4a2b846fe732ea3 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Fri, 25 Oct 2024 18:52:56 +0000 Subject: [PATCH 424/505] refactor: delete IsWellOrderLimitElement.lean (#15904) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit I understand this is quite a bold move, and I hope not to come off as confrontational. I argue that `Order/IsWellOrderLimitElement.lean` duplicates existing API, and that all the theorems proven within it already exist in some other form. First, note that the typeclass assumptions `LinearOrder α` + `IsWellOrder α (· < ·)` **very nearly** imply `ConditionallyCompleteLinearOrderBot α` through [`IsWellOrder.conditionallyCompleteLinearOrderBot`](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Order/ConditionallyCompleteLattice/Basic.html#IsWellOrder.conditionallyCompleteLinearOrderBot). The only missing assumption is `OrderBot α`, which actually follows from `Nonempty α` through [`WellFoundedLT.toOrderBot`](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Order/WellFounded.html#WellFoundedLT.toOrderBot). All theorems in this file assume the existence of some `a : α`, so that's covered. Most of the only other file importing this one assumes `OrderBot α` directly. `ConditionallyCompleteLinearOrder α` in turn implies `SuccOrder α` through `ConditionallyCompleteLinearOrder.toSuccOrder` (#15863), so assuming it yields no loss of generality. In fact, doing things this way means we can set nicer def-eqs for `@succ α`, such as `succ o = o + 1` on ordinals. As for `IsWellOrderLimitElement`, we already have [`Order.IsSuccLimit`](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Order/SuccPred/Limit.html#Order.IsSuccLimit) with only trivial typeclass assumptions. To summarize, here are the correspondences between definitions and theorems in this file and other definitions and theorems in Mathlib, applicable under equivalent typeclasss assumptions: - `wellOrderSucc` → [`Order.succ`](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Order/SuccPred/Basic.html#Order.succ) - `self_le_wellOrderSucc` → [`Order.le_succ`](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Order/SuccPred/Basic.html#Order.le_succ) - `wellOrderSucc_le` → [`Order.succ_le_of_lt`](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Order/SuccPred/Basic.html#Order.succ_le_of_lt) - `self_lt_wellOrderSucc` → [`Order.lt_succ_of_not_isMax`](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Order/SuccPred/Basic.html#Order.lt_succ_of_not_isMax) - `le_of_lt_wellOrderSucc` → [`Order.le_of_lt_succ`](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Order/SuccPred/Basic.html#Order.le_of_lt_succ). - `IsWellOrderLimitElement`→ [`Order.IsSuccLimit`](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Order/SuccPred/Limit.html#Order.IsSuccLimit) - `IsWellOrderLimitElement.neq_bot` → (already implied by `IsSuccLimit`) - `IsWellOrderLimitElement.bot_lt` → [`Ne.bot_lt`](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Order/BoundedOrder.html#Ne.bot_lt) - `IsWellOrderLimitElement.wellOrderSucc_lt` → [`Order.IsSuccLimit.succ_lt`](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Order/SuccPred/Limit.html#Order.IsSuccLimit.succ_lt) - `eq_bot_or_eq_succ_or_isWellOrderLimitElement` → [`Order.isSuccLimitRecOn`](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Order/SuccPred/Limit.html#Order.isSuccLimitRecOn) - `Nat.wellOrderSucc_eq` → [`Nat.succ_eq_succ`](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Data/Nat/SuccPred.html#Nat.succ_eq_succ) - `Nat.not_isWellOrderLimitElement` → [`Order.IsSuccLimit.isMin_of_noMax`](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Order/SuccPred/Limit.html#Order.IsSuccLimit.isMin_of_noMax) Co-authored-by: leanprover-community-mathlib4-bot --- Mathlib.lean | 1 - .../CategoryTheory/SmallObject/Iteration.lean | 62 +++++++----- Mathlib/Order/IsWellOrderLimitElement.lean | 98 ------------------- 3 files changed, 35 insertions(+), 126 deletions(-) delete mode 100644 Mathlib/Order/IsWellOrderLimitElement.lean diff --git a/Mathlib.lean b/Mathlib.lean index b0ddfe0730762..0307049fa8db8 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3801,7 +3801,6 @@ import Mathlib.Order.Interval.Set.SurjOn import Mathlib.Order.Interval.Set.UnorderedInterval import Mathlib.Order.Interval.Set.WithBotTop import Mathlib.Order.Irreducible -import Mathlib.Order.IsWellOrderLimitElement import Mathlib.Order.Iterate import Mathlib.Order.JordanHolder import Mathlib.Order.KonigLemma diff --git a/Mathlib/CategoryTheory/SmallObject/Iteration.lean b/Mathlib/CategoryTheory/SmallObject/Iteration.lean index a9470418b418d..f4d4af8823a9f 100644 --- a/Mathlib/CategoryTheory/SmallObject/Iteration.lean +++ b/Mathlib/CategoryTheory/SmallObject/Iteration.lean @@ -5,7 +5,8 @@ Authors: Joël Riou -/ import Mathlib.CategoryTheory.Category.Preorder import Mathlib.CategoryTheory.Limits.IsLimit -import Mathlib.Order.IsWellOrderLimitElement +import Mathlib.Order.ConditionallyCompleteLattice.Basic +import Mathlib.Order.SuccPred.Limit /-! # Transfinite iterations of a functor @@ -37,7 +38,7 @@ namespace CategoryTheory open Category Limits variable {C : Type*} [Category C] {Φ : C ⥤ C} (ε : 𝟭 C ⟶ Φ) - {J : Type u} [LinearOrder J] + {J : Type u} [Preorder J] namespace Functor @@ -45,11 +46,11 @@ namespace Iteration variable {j : J} (F : { i // i ≤ j } ⥤ C) -/-- The map `F.obj ⟨i, _⟩ ⟶ F.obj ⟨wellOrderSucc i, _⟩` when `F : { i // i ≤ j } ⥤ C` +/-- The map `F.obj ⟨i, _⟩ ⟶ F.obj ⟨Order.succ i, _⟩` when `F : { i // i ≤ j } ⥤ C` and `i : J` is such that `i < j`. -/ -noncomputable abbrev mapSucc' [IsWellOrder J (· < ·)] (i : J) (hi : i < j) : - F.obj ⟨i, hi.le⟩ ⟶ F.obj ⟨wellOrderSucc i, wellOrderSucc_le hi⟩ := - F.map (homOfLE (by simpa only [Subtype.mk_le_mk] using self_le_wellOrderSucc i)) +noncomputable abbrev mapSucc' [SuccOrder J] (i : J) (hi : i < j) : + F.obj ⟨i, hi.le⟩ ⟶ F.obj ⟨Order.succ i, Order.succ_le_of_lt hi⟩ := + F.map <| homOfLE <| Subtype.mk_le_mk.2 <| Order.le_succ i variable {i : J} (hi : i ≤ j) @@ -78,29 +79,26 @@ def coconeOfLE : Cocone (restrictionLT F hi) where end Iteration -variable [IsWellOrder J (· < ·)] [OrderBot J] - /-- The category of `j`th iterations of a functor `Φ` equipped with a natural transformation `ε : 𝟭 C ⟶ Φ`. An object consists of the data of all iterations of `Φ` for `i : J` such that `i ≤ j` (this is the field `F`). Such objects are equipped with data and properties which characterizes the iterations up to a unique isomorphism for the three types of elements: `⊥`, successors, limit elements. -/ -structure Iteration (j : J) where +structure Iteration [Preorder J] [OrderBot J] [SuccOrder J] (j : J) where /-- The data of all `i`th iterations for `i : J` such that `i ≤ j`. -/ F : { i // i ≤ j } ⥤ C ⥤ C /-- The zeroth iteration is the identity functor. -/ isoZero : F.obj ⟨⊥, bot_le⟩ ≅ 𝟭 C /-- The iteration on a successor element is obtained by composition of the previous iteration with `Φ`. -/ - isoSucc (i : J) (hi : i < j) : - F.obj ⟨wellOrderSucc i, wellOrderSucc_le hi⟩ ≅ F.obj ⟨i, hi.le⟩ ⋙ Φ + isoSucc (i : J) (hi : i < j) : F.obj ⟨Order.succ i, Order.succ_le_of_lt hi⟩ ≅ F.obj ⟨i, hi.le⟩ ⋙ Φ /-- The natural map from an iteration to its successor is induced by `ε`. -/ mapSucc'_eq (i : J) (hi : i < j) : Iteration.mapSucc' F i hi = whiskerLeft _ ε ≫ (isoSucc i hi).inv /-- If `i` is a limit element, the `i`th iteration is the colimit of `k`th iterations for `k < i`. -/ - isColimit (i : J) [IsWellOrderLimitElement i] (hi : i ≤ j) : - IsColimit (Iteration.coconeOfLE F hi) + isColimit (i : J) (hi : Order.IsSuccLimit i) (hij : i ≤ j) : + IsColimit (Iteration.coconeOfLE F hij) namespace Iteration @@ -109,12 +107,12 @@ variable {j : J} section -variable (iter : Φ.Iteration ε j) +variable [OrderBot J] [SuccOrder J] (iter : Φ.Iteration ε j) /-- For `iter : Φ.Iteration.ε j`, this is the map -`iter.F.obj ⟨i, _⟩ ⟶ iter.F.obj ⟨wellOrderSucc i, _⟩` if `i : J` is such that `i < j`. -/ +`iter.F.obj ⟨i, _⟩ ⟶ iter.F.obj ⟨Order.succ i, _⟩` if `i : J` is such that `i < j`. -/ noncomputable abbrev mapSucc (i : J) (hi : i < j) : - iter.F.obj ⟨i, hi.le⟩ ⟶ iter.F.obj ⟨wellOrderSucc i, wellOrderSucc_le hi⟩ := + iter.F.obj ⟨i, hi.le⟩ ⟶ iter.F.obj ⟨Order.succ i, Order.succ_le_of_lt hi⟩ := mapSucc' iter.F i hi lemma mapSucc_eq (i : J) (hi : i < j) : @@ -123,7 +121,7 @@ lemma mapSucc_eq (i : J) (hi : i < j) : end -variable (iter₁ iter₂ : Φ.Iteration ε j) +variable [OrderBot J] [SuccOrder J] (iter₁ iter₂ : Φ.Iteration ε j) /-- A morphism between two objects `iter₁` and `iter₂` in the category `Φ.Iteration ε j` of `j`th iterations of a functor `Φ` @@ -136,7 +134,7 @@ structure Hom where natTrans_app_zero : natTrans.app ⟨⊥, bot_le⟩ = iter₁.isoZero.hom ≫ iter₂.isoZero.inv := by aesop_cat natTrans_app_succ (i : J) (hi : i < j) : - natTrans.app ⟨wellOrderSucc i, wellOrderSucc_le hi⟩ = (iter₁.isoSucc i hi).hom ≫ + natTrans.app ⟨Order.succ i, Order.succ_le_of_lt hi⟩ = (iter₁.isoSucc i hi).hom ≫ whiskerRight (natTrans.app ⟨i, hi.le⟩) _ ≫ (iter₂.isoSucc i hi).inv := by aesop_cat namespace Hom @@ -172,17 +170,27 @@ instance : Category (Iteration ε j) where id := id comp := comp -instance : Subsingleton (iter₁ ⟶ iter₂) where - allEq f g := ext' (by - let P := fun (i : J) => ∀ (hi : i ≤ j), f.natTrans.app ⟨i, hi⟩ = g.natTrans.app ⟨i, hi⟩ - suffices ∀ (i : J), P i by +instance {J} {j : J} [PartialOrder J] [OrderBot J] [WellFoundedLT J] [SuccOrder J] + {iter₁ iter₂ : Iteration ε j} : + Subsingleton (iter₁ ⟶ iter₂) where + allEq f g := by + apply ext' + suffices ∀ i hi, f.natTrans.app ⟨i, hi⟩ = g.natTrans.app ⟨i, hi⟩ by ext ⟨i, hi⟩ : 2 apply this - refine fun _ => WellFoundedLT.induction _ (fun i hi hi' => ?_) - obtain rfl|⟨i, rfl, hi''⟩|_ := eq_bot_or_eq_succ_or_isWellOrderLimitElement i - · simp only [natTrans_app_zero] - · simp only [Hom.natTrans_app_succ _ i (lt_of_lt_of_le hi'' hi'), hi i hi''] - · exact (iter₁.isColimit i hi').hom_ext (fun ⟨k, hk⟩ => by simp [hi k hk])) + intro i + induction i using SuccOrder.limitRecOn with + | hm j H => + have := H.eq_bot + subst this + simp [natTrans_app_zero] + | hs j H IH => + intro hj + simp [Hom.natTrans_app_succ, IH, (Order.lt_succ_of_not_isMax H).trans_le hj] + | hl j H IH => + apply fun hj ↦ (iter₁.isColimit j H hj).hom_ext ?_ + rintro ⟨k, hk⟩ + simp [IH k hk] end Hom diff --git a/Mathlib/Order/IsWellOrderLimitElement.lean b/Mathlib/Order/IsWellOrderLimitElement.lean deleted file mode 100644 index 2d6bde7b9282b..0000000000000 --- a/Mathlib/Order/IsWellOrderLimitElement.lean +++ /dev/null @@ -1,98 +0,0 @@ -/- -Copyright (c) 2024 Joël Riou. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Joël Riou --/ -import Mathlib.Algebra.Order.Group.Nat -import Mathlib.Order.WellFounded - -/-! -# Limit elements in well-ordered types - -This file introduces two main definitions: -- `wellOrderSucc a`: the successor of an element `a` in a well-ordered type -- the typeclass `IsWellOrderLimitElement a` which asserts that an element `a` (in a -well-ordered type) is neither a successor nor the smallest element, i.e. `a` is a limit element - -Then, the lemma `eq_bot_or_eq_succ_or_isWellOrderLimitElement` shows that an element -in a well-ordered type is either `⊥`, a successor, or a limit element. - --/ - -variable {α : Type*} [LinearOrder α] - -section -variable [IsWellOrder α (· < ·)] - -/-- Given an element `a : α` in a well ordered set, this is the successor of `a`, -i.e. the smallest element strictly greater than `a` if it exists (or `a` itself otherwise). -/ -noncomputable def wellOrderSucc (a : α) : α := - (IsWellFounded.wf (r := (· < ·))).succ a - -lemma self_le_wellOrderSucc (a : α) : a ≤ wellOrderSucc a := by - by_cases h : ∃ b, a < b - · exact (IsWellFounded.wf.lt_succ h).le - · dsimp [wellOrderSucc, WellFounded.succ] - rw [dif_neg h] - -lemma wellOrderSucc_le {a b : α} (ha : a < b) : wellOrderSucc a ≤ b := by - dsimp [wellOrderSucc, WellFounded.succ] - rw [dif_pos ⟨_, ha⟩] - exact WellFounded.min_le _ ha - -lemma self_lt_wellOrderSucc {a b : α} (h : a < b) : a < wellOrderSucc a := - IsWellFounded.wf.lt_succ ⟨b, h⟩ - -lemma le_of_lt_wellOrderSucc {a b : α} (h : a < wellOrderSucc b) : a ≤ b := by - by_contra! - simpa using lt_of_lt_of_le h (wellOrderSucc_le this) - -/-- This property of an element `a : α` in a well-ordered type holds iff `a` is a -limit element, i.e. it is not a successor and it is not the smallest element of `α`. -/ -class IsWellOrderLimitElement (a : α) : Prop where - not_bot : ∃ (b : α), b < a - not_succ (b : α) (hb : b < a) : ∃ (c : α), b < c ∧ c < a - -end - -variable (a : α) [ha : IsWellOrderLimitElement a] - -lemma IsWellOrderLimitElement.neq_bot [OrderBot α] : a ≠ ⊥ := by - rintro rfl - obtain ⟨b, hb⟩ := ha.not_bot - simp at hb - -lemma IsWellOrderLimitElement.bot_lt [OrderBot α] : ⊥ < a := by - obtain h|h := eq_or_lt_of_le (@bot_le _ _ _ a) - · exact (IsWellOrderLimitElement.neq_bot a h.symm).elim - · exact h - -variable {a} -variable [IsWellOrder α (· < ·)] - -lemma IsWellOrderLimitElement.wellOrderSucc_lt {b : α} (hb : b < a) : - wellOrderSucc b < a := by - obtain ⟨c, hc₁, hc₂⟩ := ha.not_succ b hb - exact lt_of_le_of_lt (wellOrderSucc_le hc₁) hc₂ - -lemma eq_bot_or_eq_succ_or_isWellOrderLimitElement [OrderBot α] (a : α) : - a = ⊥ ∨ (∃ b, a = wellOrderSucc b ∧ b < a) ∨ IsWellOrderLimitElement a := by - refine or_iff_not_imp_left.2 <| fun h₁ ↦ or_iff_not_imp_left.2 <| fun h₂ ↦ ?_ - refine (IsWellOrderLimitElement.mk ⟨⊥, Ne.bot_lt h₁⟩ fun b hb ↦ ?_) - obtain rfl | h₃ := eq_or_lt_of_le (wellOrderSucc_le hb) - · exact (h₂ ⟨b, rfl, hb⟩).elim - · exact ⟨wellOrderSucc b, self_lt_wellOrderSucc hb, h₃⟩ - -lemma IsWellOrderLimitElement.neq_succ (a : α) (ha : a < wellOrderSucc a) - [IsWellOrderLimitElement (wellOrderSucc a)] : False := by - simpa using IsWellOrderLimitElement.wellOrderSucc_lt ha - -@[simp] -lemma Nat.wellOrderSucc_eq (a : ℕ) : wellOrderSucc a = succ a := - le_antisymm (wellOrderSucc_le (Nat.lt_succ_self a)) - (Nat.succ_le.1 (self_lt_wellOrderSucc (Nat.lt_succ_self a))) - -lemma Nat.not_isWellOrderLimitElement (a : ℕ) [IsWellOrderLimitElement a] : False := by - obtain _|a := a - · simpa using IsWellOrderLimitElement.neq_bot (0 : ℕ) - · simpa using IsWellOrderLimitElement.wellOrderSucc_lt (Nat.lt_succ_self a) From 2a1cffe71f425945911deaebfd02a990443fbb24 Mon Sep 17 00:00:00 2001 From: Christopher Hoskin Date: Fri, 25 Oct 2024 19:23:21 +0000 Subject: [PATCH 425/505] feature(LinearAlgebra/QuadraticForm/Basis): toBilinHom (#18212) For each basis `bm`, `QuadraticMap.toBilinHom` is the linear map that sends a quadratic map on a module `M` over `R` to its associated symmetric bilinear map. For free modules, `toBilinHom` is the `toBilin` equivalent of `QuadraticMap.associatedHom` Co-authored-by: Eric Wieser Co-authored-by: Christopher Hoskin --- .../LinearAlgebra/QuadraticForm/Basis.lean | 30 +++++++++++++++++++ 1 file changed, 30 insertions(+) diff --git a/Mathlib/LinearAlgebra/QuadraticForm/Basis.lean b/Mathlib/LinearAlgebra/QuadraticForm/Basis.lean index c2c99925d7bb1..40f483a2ba81e 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/Basis.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/Basis.lean @@ -13,6 +13,8 @@ does not require `Invertible (2 : R)`. Unlike that definition, this only works i a basis. -/ +open LinearMap (BilinMap) + namespace QuadraticMap variable {ι R M N} [LinearOrder ι] @@ -60,4 +62,32 @@ theorem _root_.LinearMap.BilinMap.toQuadraticMap_surjective [Module.Free R M] : letI : LinearOrder ι := IsWellOrder.linearOrder WellOrderingRel exact ⟨_, toQuadraticMap_toBilin _ b⟩ +@[simp] +lemma add_toBilin (bm : Basis ι R M) (Q₁ Q₂ : QuadraticMap R M N) : + (Q₁ + Q₂).toBilin bm = Q₁.toBilin bm + Q₂.toBilin bm := by + refine bm.ext fun i => bm.ext fun j => ?_ + obtain h | rfl | h := lt_trichotomy i j + · simp [h.ne, h, toBilin_apply, polar_add] + · simp [toBilin_apply] + · simp [h.ne', h.not_lt, toBilin_apply, polar_add] + +variable (S) [CommSemiring S] [Algebra S R] +variable [Module S N] [IsScalarTower S R N] + +@[simp] +lemma smul_toBilin (bm : Basis ι R M) (s : S) (Q : QuadraticMap R M N) : + (s • Q).toBilin bm = s • Q.toBilin bm := by + refine bm.ext fun i => bm.ext fun j => ?_ + obtain h | rfl | h := lt_trichotomy i j + · simp [h.ne, h, toBilin_apply, polar_smul] + · simp [toBilin_apply] + · simp [h.ne', h.not_lt, toBilin_apply] + +/-- `QuadraticMap.toBilin` as an S-linear map -/ +@[simps] +noncomputable def toBilinHom (bm : Basis ι R M) : QuadraticMap R M N →ₗ[S] BilinMap R M N where + toFun Q := Q.toBilin bm + map_add' := add_toBilin bm + map_smul' := smul_toBilin S bm + end QuadraticMap From eba9c377332bae1a588c85fc07a1a4012ea29299 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Fri, 25 Oct 2024 20:03:54 +0000 Subject: [PATCH 426/505] chore(Data/Nat/Prime): split `Defs`; rearrange `Basic` (#18233) This PR reduces the theory included with `Data/Nat/Prime/Defs.lean`, in particular allowing us to define `CharP` without needing theory on ordered monoids. Most of the contents got moved to `Basic.lean`, which would increase the downstream imports for quite a few file. So in turn I split up `Basic.lean` based on the imports and logical structure: * Results related to the factorial go to `Prime/Factorial.lean` * Results related to integers go to `Prime/Int.lean` (which should probably be merged with `Data/Int/NatPrime.lean) * The proof that there are infinitely many primes goes to `Prime/Infinite.lean` * Some results on powers of prime that needed a higher amount of ordered monoid theory go to `Prime/Pow.lean` This is not a universal import win, since some files that imported `Defs.lean` now need to import `Basic.lean` instead, but I think that's worth it for keeping `Defs` files succinct. --- Archive/Imo/Imo2019Q4.lean | 2 +- Mathlib.lean | 4 + Mathlib/Algebra/Order/Floor/Prime.lean | 2 +- Mathlib/Data/Int/NatPrime.lean | 1 + Mathlib/Data/Nat/Choose/Dvd.lean | 2 +- Mathlib/Data/Nat/Factorization/PrimePow.lean | 1 + Mathlib/Data/Nat/Factors.lean | 2 +- Mathlib/Data/Nat/Prime/Basic.lean | 105 +++++------------- Mathlib/Data/Nat/Prime/Defs.lean | 28 +---- Mathlib/Data/Nat/Prime/Factorial.lean | 28 +++++ Mathlib/Data/Nat/Prime/Infinite.lean | 49 ++++++++ Mathlib/Data/Nat/Prime/Int.lean | 47 ++++++++ Mathlib/Data/Nat/Prime/Pow.lean | 31 ++++++ Mathlib/Data/Nat/PrimeFin.lean | 2 +- Mathlib/Dynamics/PeriodicPts.lean | 1 + .../LegendreSymbol/GaussEisensteinLemmas.lean | 3 +- Mathlib/NumberTheory/Multiplicity.lean | 2 +- .../NumberTheory/Padics/PadicVal/Basic.lean | 1 + Mathlib/NumberTheory/Primorial.lean | 2 +- Mathlib/RingTheory/Int/Basic.lean | 4 +- Mathlib/Tactic/NormNum/Prime.lean | 2 +- test/observe.lean | 2 +- 22 files changed, 206 insertions(+), 115 deletions(-) create mode 100644 Mathlib/Data/Nat/Prime/Factorial.lean create mode 100644 Mathlib/Data/Nat/Prime/Infinite.lean create mode 100644 Mathlib/Data/Nat/Prime/Int.lean create mode 100644 Mathlib/Data/Nat/Prime/Pow.lean diff --git a/Archive/Imo/Imo2019Q4.lean b/Archive/Imo/Imo2019Q4.lean index ddf0040f14bd9..fe10780d6586f 100644 --- a/Archive/Imo/Imo2019Q4.lean +++ b/Archive/Imo/Imo2019Q4.lean @@ -5,7 +5,7 @@ Authors: Floris van Doorn -/ import Mathlib.Data.Nat.Factorial.BigOperators import Mathlib.Data.Nat.Multiplicity -import Mathlib.Data.Nat.Prime.Basic +import Mathlib.Data.Nat.Prime.Int import Mathlib.Tactic.IntervalCases import Mathlib.Tactic.GCongr diff --git a/Mathlib.lean b/Mathlib.lean index 0307049fa8db8..61df8bf2227e8 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2544,6 +2544,10 @@ import Mathlib.Data.Nat.PartENat import Mathlib.Data.Nat.Periodic import Mathlib.Data.Nat.Prime.Basic import Mathlib.Data.Nat.Prime.Defs +import Mathlib.Data.Nat.Prime.Factorial +import Mathlib.Data.Nat.Prime.Infinite +import Mathlib.Data.Nat.Prime.Int +import Mathlib.Data.Nat.Prime.Pow import Mathlib.Data.Nat.PrimeFin import Mathlib.Data.Nat.Set import Mathlib.Data.Nat.Size diff --git a/Mathlib/Algebra/Order/Floor/Prime.lean b/Mathlib/Algebra/Order/Floor/Prime.lean index f3545648ae511..506029de21811 100644 --- a/Mathlib/Algebra/Order/Floor/Prime.lean +++ b/Mathlib/Algebra/Order/Floor/Prime.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Yuyang Zhao. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yuyang Zhao -/ -import Mathlib.Data.Nat.Prime.Basic +import Mathlib.Data.Nat.Prime.Infinite import Mathlib.Topology.Algebra.Order.Floor /-! diff --git a/Mathlib/Data/Int/NatPrime.lean b/Mathlib/Data/Int/NatPrime.lean index f1cbf001f3416..ee1e51f543988 100644 --- a/Mathlib/Data/Int/NatPrime.lean +++ b/Mathlib/Data/Int/NatPrime.lean @@ -3,6 +3,7 @@ Copyright (c) 2020 Bryan Gin-ge Chen. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kevin Lacker, Bryan Gin-ge Chen -/ +import Mathlib.Algebra.Group.Int import Mathlib.Data.Nat.Prime.Basic /-! diff --git a/Mathlib/Data/Nat/Choose/Dvd.lean b/Mathlib/Data/Nat/Choose/Dvd.lean index e5cca9d05b77b..0272d236dd30a 100644 --- a/Mathlib/Data/Nat/Choose/Dvd.lean +++ b/Mathlib/Data/Nat/Choose/Dvd.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Patrick Stevens -/ import Mathlib.Data.Nat.Choose.Basic -import Mathlib.Data.Nat.Prime.Basic +import Mathlib.Data.Nat.Prime.Factorial /-! # Divisibility properties of binomial coefficients diff --git a/Mathlib/Data/Nat/Factorization/PrimePow.lean b/Mathlib/Data/Nat/Factorization/PrimePow.lean index 37333d9e42ecf..f36f7a25b7efd 100644 --- a/Mathlib/Data/Nat/Factorization/PrimePow.lean +++ b/Mathlib/Data/Nat/Factorization/PrimePow.lean @@ -5,6 +5,7 @@ Authors: Bhavik Mehta -/ import Mathlib.Algebra.IsPrimePow import Mathlib.Data.Nat.Factorization.Basic +import Mathlib.Data.Nat.Prime.Pow /-! # Prime powers and factorizations diff --git a/Mathlib/Data/Nat/Factors.lean b/Mathlib/Data/Nat/Factors.lean index c3d6f9ad26ecb..ff8e8166a2c28 100644 --- a/Mathlib/Data/Nat/Factors.lean +++ b/Mathlib/Data/Nat/Factors.lean @@ -5,7 +5,7 @@ Authors: Leonardo de Moura, Jeremy Avigad, Mario Carneiro -/ import Mathlib.Algebra.BigOperators.Ring.List import Mathlib.Data.Nat.GCD.Basic -import Mathlib.Data.Nat.Prime.Defs +import Mathlib.Data.Nat.Prime.Basic import Mathlib.Data.List.Prime import Mathlib.Data.List.Sort import Mathlib.Data.List.Perm.Subperm diff --git a/Mathlib/Data/Nat/Prime/Basic.lean b/Mathlib/Data/Nat/Prime/Basic.lean index 5a4b6a4fb3db7..a0c7d3f9fc836 100644 --- a/Mathlib/Data/Nat/Prime/Basic.lean +++ b/Mathlib/Data/Nat/Prime/Basic.lean @@ -4,22 +4,17 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Jeremy Avigad, Mario Carneiro -/ import Mathlib.Algebra.Associated.Basic -import Mathlib.Algebra.Order.Monoid.Unbundled.Pow -import Mathlib.Algebra.Ring.Int -import Mathlib.Data.Nat.Factorial.Basic +import Mathlib.Algebra.Ring.Parity import Mathlib.Data.Nat.Prime.Defs -import Mathlib.Order.Bounds.Basic /-! -## Notable Theorems +# Prime numbers -- `Nat.exists_infinite_primes`: Euclid's theorem that there exist infinitely many prime numbers. - This also appears as `Nat.not_bddAbove_setOf_prime` and `Nat.infinite_setOf_prime` (the latter - in `Data.Nat.PrimeFin`). +This file develops the theory of prime numbers: natural numbers `p ≥ 2` whose only divisors are +`p` and `1`. -/ - open Bool Subtype open Nat @@ -27,6 +22,29 @@ open Nat namespace Nat variable {n : ℕ} +theorem prime_mul_iff {a b : ℕ} : Nat.Prime (a * b) ↔ a.Prime ∧ b = 1 ∨ b.Prime ∧ a = 1 := by + simp only [irreducible_mul_iff, ← irreducible_iff_nat_prime, Nat.isUnit_iff] + +theorem not_prime_mul {a b : ℕ} (a1 : a ≠ 1) (b1 : b ≠ 1) : ¬Prime (a * b) := by + simp [prime_mul_iff, _root_.not_or, *] + +theorem not_prime_mul' {a b n : ℕ} (h : a * b = n) (h₁ : a ≠ 1) (h₂ : b ≠ 1) : ¬Prime n := + h ▸ not_prime_mul h₁ h₂ + +theorem Prime.dvd_iff_eq {p a : ℕ} (hp : p.Prime) (a1 : a ≠ 1) : a ∣ p ↔ p = a := by + refine ⟨?_, by rintro rfl; rfl⟩ + rintro ⟨j, rfl⟩ + rcases prime_mul_iff.mp hp with (⟨_, rfl⟩ | ⟨_, rfl⟩) + · exact mul_one _ + · exact (a1 rfl).elim + +theorem Prime.eq_two_or_odd {p : ℕ} (hp : Prime p) : p = 2 ∨ p % 2 = 1 := + p.mod_two_eq_zero_or_one.imp_left fun h => + ((hp.eq_one_or_self_of_dvd 2 (dvd_of_mod_eq_zero h)).resolve_left (by decide)).symm + +theorem Prime.eq_two_or_odd' {p : ℕ} (hp : Prime p) : p = 2 ∨ Odd p := + Or.imp_right (fun h => ⟨p / 2, (div_add_mod p 2).symm.trans (congr_arg _ h)⟩) hp.eq_two_or_odd + section theorem Prime.five_le_of_ne_two_of_ne_three {p : ℕ} (hp : p.Prime) (h_two : p ≠ 2) @@ -78,26 +96,6 @@ theorem dvd_of_forall_prime_mul_dvd {a b : ℕ} obtain ⟨p, hp⟩ := exists_prime_and_dvd ha exact _root_.trans (dvd_mul_left a p) (hdvd p hp.1 hp.2) -/-- Euclid's theorem on the **infinitude of primes**. -Here given in the form: for every `n`, there exists a prime number `p ≥ n`. -/ -theorem exists_infinite_primes (n : ℕ) : ∃ p, n ≤ p ∧ Prime p := - let p := minFac (n ! + 1) - have f1 : n ! + 1 ≠ 1 := ne_of_gt <| succ_lt_succ <| factorial_pos _ - have pp : Prime p := minFac_prime f1 - have np : n ≤ p := - le_of_not_ge fun h => - have h₁ : p ∣ n ! := dvd_factorial (minFac_pos _) h - have h₂ : p ∣ 1 := (Nat.dvd_add_iff_right h₁).2 (minFac_dvd _) - pp.not_dvd_one h₂ - ⟨p, np, pp⟩ - -/-- A version of `Nat.exists_infinite_primes` using the `BddAbove` predicate. -/ -theorem not_bddAbove_setOf_prime : ¬BddAbove { p | Prime p } := by - rw [not_bddAbove_iff] - intro n - obtain ⟨p, hi, hp⟩ := exists_infinite_primes n.succ - exact ⟨p, hp, hi⟩ - theorem Prime.even_iff {p : ℕ} (hp : Prime p) : Even p ↔ p = 2 := by rw [even_iff_two_dvd, prime_dvd_prime_iff_eq prime_two hp, eq_comm] @@ -159,18 +157,6 @@ theorem Prime.pow_eq_iff {p a k : ℕ} (hp : p.Prime) : a ^ k = p ↔ a = p ∧ rw [← h] at hp rw [← h, hp.eq_one_of_pow, eq_self_iff_true, _root_.and_true, pow_one] -theorem pow_minFac {n k : ℕ} (hk : k ≠ 0) : (n ^ k).minFac = n.minFac := by - rcases eq_or_ne n 1 with (rfl | hn) - · simp - have hnk : n ^ k ≠ 1 := fun hk' => hn ((pow_eq_one_iff hk).1 hk') - apply (minFac_le_of_dvd (minFac_prime hn).two_le ((minFac_dvd n).pow hk)).antisymm - apply - minFac_le_of_dvd (minFac_prime hnk).two_le - ((minFac_prime hnk).dvd_of_dvd_pow (minFac_dvd _)) - -theorem Prime.pow_minFac {p k : ℕ} (hp : p.Prime) (hk : k ≠ 0) : (p ^ k).minFac = p := by - rw [Nat.pow_minFac hk, hp.minFac_eq] - theorem Prime.mul_eq_prime_sq_iff {x y p : ℕ} (hp : p.Prime) (hx : x ≠ 1) (hy : y ≠ 1) : x * y = p ^ 2 ↔ x = p ∧ y = p := by refine ⟨fun h => ?_, fun ⟨h₁, h₂⟩ => h₁.symm ▸ h₂.symm ▸ (sq _).symm⟩ @@ -197,14 +183,6 @@ theorem Prime.mul_eq_prime_sq_iff {x y p : ℕ} (hp : p.Prime) (hx : x ≠ 1) (h rw [sq, Nat.mul_right_eq_self_iff (Nat.mul_pos hp.pos hp.pos : 0 < a * a)] at h exact h -theorem Prime.dvd_factorial : ∀ {n p : ℕ} (_ : Prime p), p ∣ n ! ↔ p ≤ n - | 0, _, hp => iff_of_false hp.not_dvd_one (not_le_of_lt hp.pos) - | n + 1, p, hp => by - rw [factorial_succ, hp.dvd_mul, Prime.dvd_factorial hp] - exact - ⟨fun h => h.elim (le_of_dvd (succ_pos _)) le_succ_of_le, fun h => - (_root_.lt_or_eq_of_le h).elim (Or.inr ∘ le_of_lt_succ) fun h => Or.inl <| by rw [h]⟩ - theorem Prime.coprime_pow_of_not_dvd {p m a : ℕ} (pp : Prime p) (h : ¬p ∣ a) : Coprime a (p ^ m) := (pp.coprime_iff_not_dvd.2 h).symm.pow_right _ @@ -269,33 +247,4 @@ theorem succ_dvd_or_succ_dvd_of_succ_sum_dvd_mul {p : ℕ} (p_prime : Prime p) { exact hpd5.elim (fun h : p ∣ m / p ^ k => Or.inl <| mul_dvd_of_dvd_div hpm h) fun h : p ∣ n / p ^ l => Or.inr <| mul_dvd_of_dvd_div hpn h -theorem prime_iff_prime_int {p : ℕ} : p.Prime ↔ _root_.Prime (p : ℤ) := - ⟨fun hp => - ⟨Int.natCast_ne_zero_iff_pos.2 hp.pos, mt Int.isUnit_iff_natAbs_eq.1 hp.ne_one, fun a b h => by - rw [← Int.dvd_natAbs, Int.natCast_dvd_natCast, Int.natAbs_mul, hp.dvd_mul] at h - rwa [← Int.dvd_natAbs, Int.natCast_dvd_natCast, ← Int.dvd_natAbs, Int.natCast_dvd_natCast]⟩, - fun hp => - Nat.prime_iff.2 - ⟨Int.natCast_ne_zero.1 hp.1, - (mt Nat.isUnit_iff.1) fun h => by simp [h, not_prime_one] at hp, fun a b => by - simpa only [Int.natCast_dvd_natCast, (Int.ofNat_mul _ _).symm] using hp.2.2 a b⟩⟩ - -/-- Two prime powers with positive exponents are equal only when the primes and the -exponents are equal. -/ -lemma Prime.pow_inj {p q m n : ℕ} (hp : p.Prime) (hq : q.Prime) - (h : p ^ (m + 1) = q ^ (n + 1)) : p = q ∧ m = n := by - have H := dvd_antisymm (Prime.dvd_of_dvd_pow hp <| h ▸ dvd_pow_self p (succ_ne_zero m)) - (Prime.dvd_of_dvd_pow hq <| h.symm ▸ dvd_pow_self q (succ_ne_zero n)) - exact ⟨H, succ_inj'.mp <| Nat.pow_right_injective hq.two_le (H ▸ h)⟩ - end Nat - -namespace Int - -theorem prime_two : Prime (2 : ℤ) := - Nat.prime_iff_prime_int.mp Nat.prime_two - -theorem prime_three : Prime (3 : ℤ) := - Nat.prime_iff_prime_int.mp Nat.prime_three - -end Int diff --git a/Mathlib/Data/Nat/Prime/Defs.lean b/Mathlib/Data/Nat/Prime/Defs.lean index 1741083417a66..e6a3e5bd1a467 100644 --- a/Mathlib/Data/Nat/Prime/Defs.lean +++ b/Mathlib/Data/Nat/Prime/Defs.lean @@ -4,8 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Jeremy Avigad, Mario Carneiro -/ import Batteries.Data.Nat.Gcd -import Mathlib.Algebra.Prime.Lemmas -import Mathlib.Algebra.Ring.Parity +import Mathlib.Algebra.Prime.Defs +import Mathlib.Algebra.Ring.Nat +import Mathlib.Order.Lattice /-! # Prime numbers @@ -161,29 +162,6 @@ theorem prime_dvd_prime_iff_eq {p q : ℕ} (pp : p.Prime) (qp : q.Prime) : p ∣ theorem Prime.not_dvd_one {p : ℕ} (pp : Prime p) : ¬p ∣ 1 := Irreducible.not_dvd_one pp -theorem prime_mul_iff {a b : ℕ} : Nat.Prime (a * b) ↔ a.Prime ∧ b = 1 ∨ b.Prime ∧ a = 1 := by - simp only [irreducible_mul_iff, ← irreducible_iff_nat_prime, Nat.isUnit_iff] - -theorem not_prime_mul {a b : ℕ} (a1 : a ≠ 1) (b1 : b ≠ 1) : ¬Prime (a * b) := by - simp [prime_mul_iff, _root_.not_or, *] - -theorem not_prime_mul' {a b n : ℕ} (h : a * b = n) (h₁ : a ≠ 1) (h₂ : b ≠ 1) : ¬Prime n := - h ▸ not_prime_mul h₁ h₂ - -theorem Prime.dvd_iff_eq {p a : ℕ} (hp : p.Prime) (a1 : a ≠ 1) : a ∣ p ↔ p = a := by - refine ⟨?_, by rintro rfl; rfl⟩ - rintro ⟨j, rfl⟩ - rcases prime_mul_iff.mp hp with (⟨_, rfl⟩ | ⟨_, rfl⟩) - · exact mul_one _ - · exact (a1 rfl).elim - -theorem Prime.eq_two_or_odd {p : ℕ} (hp : Prime p) : p = 2 ∨ p % 2 = 1 := - p.mod_two_eq_zero_or_one.imp_left fun h => - ((hp.eq_one_or_self_of_dvd 2 (dvd_of_mod_eq_zero h)).resolve_left (by decide)).symm - -theorem Prime.eq_two_or_odd' {p : ℕ} (hp : Prime p) : p = 2 ∨ Odd p := - Or.imp_right (fun h => ⟨p / 2, (div_add_mod p 2).symm.trans (congr_arg _ h)⟩) hp.eq_two_or_odd - section MinFac theorem minFac_lemma (n k : ℕ) (h : ¬n < k * k) : sqrt n - k < sqrt n + 2 - k := diff --git a/Mathlib/Data/Nat/Prime/Factorial.lean b/Mathlib/Data/Nat/Prime/Factorial.lean new file mode 100644 index 0000000000000..4c2f70a84a4c9 --- /dev/null +++ b/Mathlib/Data/Nat/Prime/Factorial.lean @@ -0,0 +1,28 @@ +/- +Copyright (c) 2015 Microsoft Corporation. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Leonardo de Moura, Jeremy Avigad, Mario Carneiro +-/ +import Mathlib.Data.Nat.Factorial.Basic +import Mathlib.Data.Nat.Prime.Defs + +/-! +# Prime natural numbers and the factorial operator + +-/ + +open Bool Subtype + +open Nat + +namespace Nat + +theorem Prime.dvd_factorial : ∀ {n p : ℕ} (_ : Prime p), p ∣ n ! ↔ p ≤ n + | 0, _, hp => iff_of_false hp.not_dvd_one (not_le_of_lt hp.pos) + | n + 1, p, hp => by + rw [factorial_succ, hp.dvd_mul, Prime.dvd_factorial hp] + exact + ⟨fun h => h.elim (le_of_dvd (succ_pos _)) le_succ_of_le, fun h => + (_root_.lt_or_eq_of_le h).elim (Or.inr ∘ le_of_lt_succ) fun h => Or.inl <| by rw [h]⟩ + +end Nat diff --git a/Mathlib/Data/Nat/Prime/Infinite.lean b/Mathlib/Data/Nat/Prime/Infinite.lean new file mode 100644 index 0000000000000..f67b58631e769 --- /dev/null +++ b/Mathlib/Data/Nat/Prime/Infinite.lean @@ -0,0 +1,49 @@ +/- +Copyright (c) 2015 Microsoft Corporation. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Leonardo de Moura, Jeremy Avigad, Mario Carneiro +-/ +import Mathlib.Data.Nat.Factorial.Basic +import Mathlib.Data.Nat.Prime.Defs +import Mathlib.Order.Bounds.Basic + +/-! +## Notable Theorems + +- `Nat.exists_infinite_primes`: Euclid's theorem that there exist infinitely many prime numbers. + This also appears as `Nat.not_bddAbove_setOf_prime` and `Nat.infinite_setOf_prime` (the latter + in `Data.Nat.PrimeFin`). + +-/ + +open Bool Subtype + +open Nat + +namespace Nat + +section Infinite + +/-- Euclid's theorem on the **infinitude of primes**. +Here given in the form: for every `n`, there exists a prime number `p ≥ n`. -/ +theorem exists_infinite_primes (n : ℕ) : ∃ p, n ≤ p ∧ Prime p := + let p := minFac (n ! + 1) + have f1 : n ! + 1 ≠ 1 := ne_of_gt <| succ_lt_succ <| factorial_pos _ + have pp : Prime p := minFac_prime f1 + have np : n ≤ p := + le_of_not_ge fun h => + have h₁ : p ∣ n ! := dvd_factorial (minFac_pos _) h + have h₂ : p ∣ 1 := (Nat.dvd_add_iff_right h₁).2 (minFac_dvd _) + pp.not_dvd_one h₂ + ⟨p, np, pp⟩ + +/-- A version of `Nat.exists_infinite_primes` using the `BddAbove` predicate. -/ +theorem not_bddAbove_setOf_prime : ¬BddAbove { p | Prime p } := by + rw [not_bddAbove_iff] + intro n + obtain ⟨p, hi, hp⟩ := exists_infinite_primes n.succ + exact ⟨p, hp, hi⟩ + +end Infinite + +end Nat diff --git a/Mathlib/Data/Nat/Prime/Int.lean b/Mathlib/Data/Nat/Prime/Int.lean new file mode 100644 index 0000000000000..c5f0347d3b6c8 --- /dev/null +++ b/Mathlib/Data/Nat/Prime/Int.lean @@ -0,0 +1,47 @@ +/- +Copyright (c) 2015 Microsoft Corporation. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Leonardo de Moura, Jeremy Avigad, Mario Carneiro +-/ +import Mathlib.Algebra.Ring.Int +import Mathlib.Data.Nat.Prime.Basic + +/-! +# Prime numbers in the naturals and the integers + +TODO: This file can probably be merged with `Data/Nat/Int/NatPrime.lean`. +-/ + + +namespace Nat + +theorem prime_iff_prime_int {p : ℕ} : p.Prime ↔ _root_.Prime (p : ℤ) := + ⟨fun hp => + ⟨Int.natCast_ne_zero_iff_pos.2 hp.pos, mt Int.isUnit_iff_natAbs_eq.1 hp.ne_one, fun a b h => by + rw [← Int.dvd_natAbs, Int.natCast_dvd_natCast, Int.natAbs_mul, hp.dvd_mul] at h + rwa [← Int.dvd_natAbs, Int.natCast_dvd_natCast, ← Int.dvd_natAbs, Int.natCast_dvd_natCast]⟩, + fun hp => + Nat.prime_iff.2 + ⟨Int.natCast_ne_zero.1 hp.1, + (mt Nat.isUnit_iff.1) fun h => by simp [h, not_prime_one] at hp, fun a b => by + simpa only [Int.natCast_dvd_natCast, (Int.ofNat_mul _ _).symm] using hp.2.2 a b⟩⟩ + +/-- Two prime powers with positive exponents are equal only when the primes and the +exponents are equal. -/ +lemma Prime.pow_inj {p q m n : ℕ} (hp : p.Prime) (hq : q.Prime) + (h : p ^ (m + 1) = q ^ (n + 1)) : p = q ∧ m = n := by + have H := dvd_antisymm (Prime.dvd_of_dvd_pow hp <| h ▸ dvd_pow_self p (succ_ne_zero m)) + (Prime.dvd_of_dvd_pow hq <| h.symm ▸ dvd_pow_self q (succ_ne_zero n)) + exact ⟨H, succ_inj'.mp <| Nat.pow_right_injective hq.two_le (H ▸ h)⟩ + +end Nat + +namespace Int + +theorem prime_two : Prime (2 : ℤ) := + Nat.prime_iff_prime_int.mp Nat.prime_two + +theorem prime_three : Prime (3 : ℤ) := + Nat.prime_iff_prime_int.mp Nat.prime_three + +end Int diff --git a/Mathlib/Data/Nat/Prime/Pow.lean b/Mathlib/Data/Nat/Prime/Pow.lean new file mode 100644 index 0000000000000..a103a434d4ff6 --- /dev/null +++ b/Mathlib/Data/Nat/Prime/Pow.lean @@ -0,0 +1,31 @@ +/- +Copyright (c) 2015 Microsoft Corporation. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Leonardo de Moura, Jeremy Avigad, Mario Carneiro +-/ +import Mathlib.Algebra.Order.Monoid.Unbundled.Pow +import Mathlib.Data.Nat.Prime.Basic + +/-! +# Prime numbers + +This file develops the theory of prime numbers: natural numbers `p ≥ 2` whose only divisors are +`p` and `1`. + +-/ + +namespace Nat + +theorem pow_minFac {n k : ℕ} (hk : k ≠ 0) : (n ^ k).minFac = n.minFac := by + rcases eq_or_ne n 1 with (rfl | hn) + · simp + have hnk : n ^ k ≠ 1 := fun hk' => hn ((pow_eq_one_iff hk).1 hk') + apply (minFac_le_of_dvd (minFac_prime hn).two_le ((minFac_dvd n).pow hk)).antisymm + apply + minFac_le_of_dvd (minFac_prime hnk).two_le + ((minFac_prime hnk).dvd_of_dvd_pow (minFac_dvd _)) + +theorem Prime.pow_minFac {p k : ℕ} (hp : p.Prime) (hk : k ≠ 0) : (p ^ k).minFac = p := by + rw [Nat.pow_minFac hk, hp.minFac_eq] + +end Nat diff --git a/Mathlib/Data/Nat/PrimeFin.lean b/Mathlib/Data/Nat/PrimeFin.lean index 8f1fe7411bc75..ecb04ce9d0c0a 100644 --- a/Mathlib/Data/Nat/PrimeFin.lean +++ b/Mathlib/Data/Nat/PrimeFin.lean @@ -6,7 +6,7 @@ Authors: Leonardo de Moura, Jeremy Avigad, Mario Carneiro import Mathlib.Data.Countable.Defs import Mathlib.Data.Nat.Factors import Mathlib.Data.Set.Finite -import Mathlib.Data.Nat.Prime.Basic +import Mathlib.Data.Nat.Prime.Infinite /-! # Prime numbers diff --git a/Mathlib/Dynamics/PeriodicPts.lean b/Mathlib/Dynamics/PeriodicPts.lean index 948033b25c274..6924fe05ed12c 100644 --- a/Mathlib/Dynamics/PeriodicPts.lean +++ b/Mathlib/Dynamics/PeriodicPts.lean @@ -5,6 +5,7 @@ Authors: Yury Kudryashov -/ import Mathlib.Algebra.GroupPower.IterateHom import Mathlib.Algebra.Ring.Divisibility.Basic +import Mathlib.Algebra.Ring.Int import Mathlib.Data.List.Cycle import Mathlib.Data.Nat.GCD.Basic import Mathlib.Data.Nat.Prime.Basic diff --git a/Mathlib/NumberTheory/LegendreSymbol/GaussEisensteinLemmas.lean b/Mathlib/NumberTheory/LegendreSymbol/GaussEisensteinLemmas.lean index 6927e41b5eeef..1faaafe97442e 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/GaussEisensteinLemmas.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/GaussEisensteinLemmas.lean @@ -3,8 +3,9 @@ Copyright (c) 2018 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes -/ -import Mathlib.NumberTheory.LegendreSymbol.Basic import Mathlib.Analysis.Normed.Field.Lemmas +import Mathlib.Data.Nat.Prime.Factorial +import Mathlib.NumberTheory.LegendreSymbol.Basic /-! # Lemmas of Gauss and Eisenstein diff --git a/Mathlib/NumberTheory/Multiplicity.lean b/Mathlib/NumberTheory/Multiplicity.lean index 119163114c7bf..cfa4de5aac58e 100644 --- a/Mathlib/NumberTheory/Multiplicity.lean +++ b/Mathlib/NumberTheory/Multiplicity.lean @@ -6,9 +6,9 @@ Authors: Tian Chen, Mantas Bakšys import Mathlib.Algebra.GeomSum import Mathlib.Algebra.Order.Ring.Basic import Mathlib.Algebra.Ring.Int +import Mathlib.Data.Nat.Prime.Int import Mathlib.NumberTheory.Padics.PadicVal.Defs import Mathlib.RingTheory.Ideal.Quotient -import Mathlib.Data.Nat.Prime.Basic /-! # Multiplicity in Number Theory diff --git a/Mathlib/NumberTheory/Padics/PadicVal/Basic.lean b/Mathlib/NumberTheory/Padics/PadicVal/Basic.lean index 6ddb55a0ed27f..d2df11023e6f1 100644 --- a/Mathlib/NumberTheory/Padics/PadicVal/Basic.lean +++ b/Mathlib/NumberTheory/Padics/PadicVal/Basic.lean @@ -7,6 +7,7 @@ import Mathlib.NumberTheory.Divisors import Mathlib.NumberTheory.Padics.PadicVal.Defs import Mathlib.Data.Nat.MaxPowDiv import Mathlib.Data.Nat.Multiplicity +import Mathlib.Data.Nat.Prime.Int /-! # `p`-adic Valuation diff --git a/Mathlib/NumberTheory/Primorial.lean b/Mathlib/NumberTheory/Primorial.lean index 7ef909c7b0799..3867a3aa0f893 100644 --- a/Mathlib/NumberTheory/Primorial.lean +++ b/Mathlib/NumberTheory/Primorial.lean @@ -8,7 +8,7 @@ import Mathlib.Algebra.Order.BigOperators.Ring.Finset import Mathlib.Algebra.Order.Ring.Abs import Mathlib.Data.Nat.Choose.Sum import Mathlib.Data.Nat.Choose.Dvd -import Mathlib.Data.Nat.Prime.Defs +import Mathlib.Data.Nat.Prime.Basic /-! # Primorial diff --git a/Mathlib/RingTheory/Int/Basic.lean b/Mathlib/RingTheory/Int/Basic.lean index 2f194c3abb9aa..be4be78dc6d35 100644 --- a/Mathlib/RingTheory/Int/Basic.lean +++ b/Mathlib/RingTheory/Int/Basic.lean @@ -5,9 +5,9 @@ Authors: Johannes Hölzl, Jens Wagemaker, Aaron Anderson -/ import Mathlib.Algebra.EuclideanDomain.Basic import Mathlib.Algebra.EuclideanDomain.Int -import Mathlib.RingTheory.PrincipalIdealDomain import Mathlib.Algebra.GCDMonoid.Nat -import Mathlib.Data.Nat.Prime.Basic +import Mathlib.Data.Nat.Prime.Int +import Mathlib.RingTheory.PrincipalIdealDomain /-! # Divisibility over ℕ and ℤ diff --git a/Mathlib/Tactic/NormNum/Prime.lean b/Mathlib/Tactic/NormNum/Prime.lean index f1e5dc961f2b4..be12f0e31793f 100644 --- a/Mathlib/Tactic/NormNum/Prime.lean +++ b/Mathlib/Tactic/NormNum/Prime.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Jeremy Avigad, Mario Carneiro -/ import Mathlib.Tactic.NormNum.Basic -import Mathlib.Data.Nat.Prime.Defs +import Mathlib.Data.Nat.Prime.Basic /-! # `norm_num` extensions on natural numbers diff --git a/test/observe.lean b/test/observe.lean index 9a2b83728aaca..b595d01117739 100644 --- a/test/observe.lean +++ b/test/observe.lean @@ -1,4 +1,4 @@ -import Mathlib.Data.Nat.Prime.Basic +import Mathlib.Data.Nat.Prime.Factorial import Mathlib.Data.Nat.Factorial.Basic open Nat From 1a5abab17a4a8835baf79e49d70edc7f0972f299 Mon Sep 17 00:00:00 2001 From: YnirPaz Date: Fri, 25 Oct 2024 20:32:11 +0000 Subject: [PATCH 427/505] feat(SetTheory/Ordinal/Topology): define Ordinal.IsClosed and Ordinal.IsAcc (#16710) Defined Ordinal.IsClosed and Ordinal.IsAcc and proved basic lemmas about them. This is closedness in the club-sense. --- Mathlib/SetTheory/Ordinal/Arithmetic.lean | 3 + Mathlib/SetTheory/Ordinal/Topology.lean | 136 ++++++++++++++++++++++ Mathlib/Topology/Basic.lean | 11 ++ 3 files changed, 150 insertions(+) diff --git a/Mathlib/SetTheory/Ordinal/Arithmetic.lean b/Mathlib/SetTheory/Ordinal/Arithmetic.lean index 6e9b1070b8f6c..865bd245f7301 100644 --- a/Mathlib/SetTheory/Ordinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Ordinal/Arithmetic.lean @@ -262,6 +262,9 @@ theorem zero_or_succ_or_limit (o : Ordinal) : o = 0 ∨ (∃ a, o = succ a) ∨ if h : ∃ a, o = succ a then Or.inr (Or.inl h) else Or.inr <| Or.inr ⟨o0, fun _a => (succ_lt_of_not_succ h).2⟩ +theorem isLimit_of_not_succ_of_ne_zero {o : Ordinal} (h : ¬∃ a, o = succ a) (h' : o ≠ 0) : + IsLimit o := ((zero_or_succ_or_limit o).resolve_left h').resolve_left h + -- TODO: this is an iff with `IsSuccPrelimit` theorem IsLimit.sSup_Iio {o : Ordinal} (h : IsLimit o) : sSup (Iio o) = o := by apply (csSup_le' (fun a ha ↦ le_of_lt ha)).antisymm diff --git a/Mathlib/SetTheory/Ordinal/Topology.lean b/Mathlib/SetTheory/Ordinal/Topology.lean index 6993d517742e4..14986ed23a289 100644 --- a/Mathlib/SetTheory/Ordinal/Topology.lean +++ b/Mathlib/SetTheory/Ordinal/Topology.lean @@ -246,4 +246,140 @@ theorem enumOrd_isNormal_iff_isClosed (hs : ¬ BddAbove s) : rw [hb] exact le_bsup.{u, u} _ _ (ha.2 _ hba) +open Set Filter Set.Notation + +/-- An ordinal is an accumulation point of a set of ordinals if it is positive and there +are elements in the set arbitrarily close to the ordinal from below. -/ +def IsAcc (o : Ordinal) (S : Set Ordinal) : Prop := + AccPt o (𝓟 S) + +/-- A set of ordinals is closed below an ordinal if it contains all of +its accumulation points below the ordinal. -/ +def IsClosedBelow (S : Set Ordinal) (o : Ordinal) : Prop := + IsClosed (Iio o ↓∩ S) + +theorem isAcc_iff (o : Ordinal) (S : Set Ordinal) : o.IsAcc S ↔ + o ≠ 0 ∧ ∀ p < o, (S ∩ Ioo p o).Nonempty := by + dsimp [IsAcc] + constructor + · rw [accPt_iff_nhds] + intro h + constructor + · rintro rfl + obtain ⟨x, hx⟩ := h (Iio 1) (Iio_mem_nhds zero_lt_one) + exact hx.2 <| lt_one_iff_zero.mp hx.1.1 + · intro p plt + obtain ⟨x, hx⟩ := h (Ioo p (o + 1)) <| Ioo_mem_nhds plt (lt_succ o) + use x + refine ⟨hx.1.2, ⟨hx.1.1.1, lt_of_le_of_ne ?_ hx.2⟩⟩ + have := hx.1.1.2 + rwa [← succ_eq_add_one, lt_succ_iff] at this + · rw [accPt_iff_nhds] + intro h u umem + obtain ⟨l, hl⟩ := exists_Ioc_subset_of_mem_nhds umem ⟨0, Ordinal.pos_iff_ne_zero.mpr h.1⟩ + obtain ⟨x, hx⟩ := h.2 l hl.1 + use x + exact ⟨⟨hl.2 ⟨hx.2.1, hx.2.2.le⟩, hx.1⟩, hx.2.2.ne⟩ + +theorem IsAcc.forall_lt {o : Ordinal} {S : Set Ordinal} (h : o.IsAcc S) : + ∀ p < o, (S ∩ Ioo p o).Nonempty := ((isAcc_iff _ _).mp h).2 + +theorem IsAcc.pos {o : Ordinal} {S : Set Ordinal} (h : o.IsAcc S) : + 0 < o := Ordinal.pos_iff_ne_zero.mpr ((isAcc_iff _ _).mp h).1 + +theorem IsAcc.isLimit {o : Ordinal} {S : Set Ordinal} (h : o.IsAcc S) : IsLimit o := by + rw [isAcc_iff] at h + refine isLimit_of_not_succ_of_ne_zero (fun ⟨x, hx⟩ ↦ ?_) h.1 + rcases h.2 x (lt_of_lt_of_le (lt_succ x) hx.symm.le) with ⟨p, hp⟩ + exact (hx.symm ▸ (succ_le_iff.mpr hp.2.1)).not_lt hp.2.2 + +theorem IsAcc.mono {o : Ordinal} {S T : Set Ordinal} (h : S ⊆ T) (ho : o.IsAcc S) : + o.IsAcc T := by + rw [isAcc_iff] at * + exact ⟨ho.1, fun p plto ↦ (ho.2 p plto).casesOn fun s hs ↦ ⟨s, h hs.1, hs.2⟩⟩ + +theorem IsAcc.inter_Ioo_nonempty {o : Ordinal} {S : Set Ordinal} (hS : o.IsAcc S) + {p : Ordinal} (hp : p < o) : (S ∩ Ioo p o).Nonempty := hS.forall_lt p hp + +-- todo: prove this for a general linear `SuccOrder`. +theorem accPt_subtype {p o : Ordinal} (S : Set Ordinal) (hpo : p < o) : + AccPt p (𝓟 S) ↔ AccPt ⟨p, hpo⟩ (𝓟 (Iio o ↓∩ S)) := by + constructor + · intro h + have plim : p.IsLimit := IsAcc.isLimit h + rw [accPt_iff_nhds] at * + intro u hu + obtain ⟨l, hl⟩ := exists_Ioc_subset_of_mem_nhds hu ⟨⟨0, plim.pos.trans hpo⟩, plim.pos⟩ + obtain ⟨x, hx⟩ := h (Ioo l (p + 1)) (Ioo_mem_nhds hl.1 (lt_add_one _)) + use ⟨x, lt_of_le_of_lt (lt_succ_iff.mp hx.1.1.2) hpo⟩ + refine ⟨?_, Subtype.coe_ne_coe.mp hx.2⟩ + exact ⟨hl.2 ⟨hx.1.1.1, by exact_mod_cast lt_succ_iff.mp hx.1.1.2⟩, hx.1.2⟩ + · intro h + rw [accPt_iff_nhds] at * + intro u hu + by_cases ho : p + 1 < o + · have ppos : p ≠ 0 := by + rintro rfl + rw [zero_add] at ho + specialize h (Iio ⟨1, ho⟩) (Iio_mem_nhds (Subtype.mk_lt_mk.mpr zero_lt_one)) + obtain ⟨_, h⟩ := h + exact h.2 <| Subtype.mk_eq_mk.mpr (lt_one_iff_zero.mp h.1.1) + have plim : p.IsLimit := by + contrapose! h + obtain ⟨q, hq⟩ := ((zero_or_succ_or_limit p).resolve_left ppos).resolve_right h + use (Ioo ⟨q, ((hq ▸ lt_succ q).trans hpo)⟩ ⟨p + 1, ho⟩) + constructor + · exact Ioo_mem_nhds (by simp only [hq, Subtype.mk_lt_mk, lt_succ]) (lt_succ p) + · intro _ mem + have aux1 := Subtype.mk_lt_mk.mp mem.1.1 + have aux2 := Subtype.mk_lt_mk.mp mem.1.2 + rw [Subtype.mk_eq_mk] + rw [hq] at aux2 ⊢ + exact ((succ_le_iff.mpr aux1).antisymm (le_of_lt_succ aux2)).symm + obtain ⟨l, hl⟩ := exists_Ioc_subset_of_mem_nhds hu ⟨0, plim.pos⟩ + obtain ⟨x, hx⟩ := h (Ioo ⟨l, hl.1.trans hpo⟩ ⟨p + 1, ho⟩) (Ioo_mem_nhds hl.1 (lt_add_one p)) + use x + exact ⟨⟨hl.2 ⟨hx.1.1.1, lt_succ_iff.mp hx.1.1.2⟩, hx.1.2⟩, fun h ↦ hx.2 (SetCoe.ext h)⟩ + have hp : o = p + 1 := (le_succ_iff_eq_or_le.mp (le_of_not_lt ho)).resolve_right + (not_le_of_lt hpo) + have ppos : p ≠ 0 := by + rintro rfl + obtain ⟨x, hx⟩ := h Set.univ univ_mem + have : ↑x < o := x.2 + simp_rw [hp, zero_add, lt_one_iff_zero] at this + exact hx.2 (SetCoe.ext this) + obtain ⟨l, hl⟩ := exists_Ioc_subset_of_mem_nhds hu ⟨0, Ordinal.pos_iff_ne_zero.mpr ppos⟩ + obtain ⟨x, hx⟩ := h (Ioi ⟨l, hl.1.trans hpo⟩) (Ioi_mem_nhds hl.1) + use x + refine ⟨⟨hl.2 ⟨hx.1.1, ?_⟩, hx.1.2⟩, fun h ↦ hx.2 (SetCoe.ext h)⟩ + rw [← lt_add_one_iff, ← hp] + exact x.2 + +theorem isClosedBelow_iff {S : Set Ordinal} {o : Ordinal} : IsClosedBelow S o ↔ + ∀ p < o, IsAcc p S → p ∈ S := by + dsimp [IsClosedBelow] + constructor + · intro h p plto hp + have : AccPt ⟨p, plto⟩ (𝓟 (Iio o ↓∩ S)) := (accPt_subtype _ _).mp hp + rw [isClosed_iff_clusterPt] at h + exact h ⟨p, plto⟩ this.clusterPt + · intro h + rw [isClosed_iff_clusterPt] + intro r hr + match clusterPt_principal.mp hr with + | .inl h => exact h + | .inr h' => exact h r.1 r.2 <| (accPt_subtype _ _).mpr h' + +alias ⟨IsClosedBelow.forall_lt, _⟩ := isClosedBelow_iff + +theorem IsClosedBelow.sInter {o : Ordinal} {S : Set (Set Ordinal)} + (h : ∀ C ∈ S, IsClosedBelow C o) : IsClosedBelow (⋂₀ S) o := by + rw [isClosedBelow_iff] + intro p plto pAcc C CmemS + exact (h C CmemS).forall_lt p plto (pAcc.mono (sInter_subset_of_mem CmemS)) + +theorem IsClosedBelow.iInter {ι : Type u} {f : ι → Set Ordinal} {o : Ordinal} + (h : ∀ i, IsClosedBelow (f i) o) : IsClosedBelow (⋂ i, f i) o := + IsClosedBelow.sInter fun _ ⟨i, hi⟩ ↦ hi ▸ (h i) + end Ordinal diff --git a/Mathlib/Topology/Basic.lean b/Mathlib/Topology/Basic.lean index b9f811270fa26..9961f73158dbd 100644 --- a/Mathlib/Topology/Basic.lean +++ b/Mathlib/Topology/Basic.lean @@ -1031,6 +1031,17 @@ theorem AccPt.mono {F G : Filter X} (h : AccPt x F) (hFG : F ≤ G) : AccPt x G theorem AccPt.clusterPt (x : X) (F : Filter X) (h : AccPt x F) : ClusterPt x F := ((acc_iff_cluster x F).mp h).mono inf_le_right +theorem clusterPt_principal {x : X} {C : Set X} : + ClusterPt x (𝓟 C) ↔ x ∈ C ∨ AccPt x (𝓟 C) := by + constructor + · intro h + by_contra! hc + rw [acc_principal_iff_cluster] at hc + simp_all only [not_false_eq_true, diff_singleton_eq_self, not_true_eq_false, hc.1] + · rintro (h | h) + · exact clusterPt_principal_iff.mpr fun _ mem ↦ ⟨x, ⟨mem_of_mem_nhds mem, h⟩⟩ + · exact h.clusterPt + /-! ### Interior, closure and frontier in terms of neighborhoods -/ From 99c058ff47a92a74da72a40723a67ca1f988d0d8 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Fri, 25 Oct 2024 20:58:32 +0000 Subject: [PATCH 428/505] chore(LinearAlgebra): use maxGenEigenspace instead of iSup over genEigenspace (#18037) --- Mathlib/Algebra/Lie/TraceForm.lean | 2 +- Mathlib/Algebra/Lie/Weights/Basic.lean | 30 ++- Mathlib/Algebra/Lie/Weights/Linear.lean | 14 +- .../InnerProductSpace/JointEigenspace.lean | 2 +- Mathlib/LinearAlgebra/Eigenspace/Basic.lean | 223 +++++++++++------- Mathlib/LinearAlgebra/Eigenspace/Pi.lean | 15 +- .../LinearAlgebra/Eigenspace/Semisimple.lean | 30 ++- .../Eigenspace/Triangularizable.lean | 151 ++++++++---- Mathlib/LinearAlgebra/Eigenspace/Zero.lean | 2 +- Mathlib/Order/SupIndep.lean | 23 ++ 10 files changed, 318 insertions(+), 174 deletions(-) diff --git a/Mathlib/Algebra/Lie/TraceForm.lean b/Mathlib/Algebra/Lie/TraceForm.lean index 62c493ca9f4c7..2d1253e2b18c9 100644 --- a/Mathlib/Algebra/Lie/TraceForm.lean +++ b/Mathlib/Algebra/Lie/TraceForm.lean @@ -277,7 +277,7 @@ 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 - · simpa only [Module.End.maxGenEigenspace_def] using IsTriangularizable.iSup_eq_top (1 ⊗ₜ[R] x) + · exact IsTriangularizable.maxGenEigenspace_eq_top (1 ⊗ₜ[R] x) · exact fun μ ↦ trace_toEnd_eq_zero_of_mem_lcs A (A ⊗[R] L) (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) diff --git a/Mathlib/Algebra/Lie/Weights/Basic.lean b/Mathlib/Algebra/Lie/Weights/Basic.lean index 1d5d4302ac426..a903c7e4c8c88 100644 --- a/Mathlib/Algebra/Lie/Weights/Basic.lean +++ b/Mathlib/Algebra/Lie/Weights/Basic.lean @@ -165,7 +165,7 @@ theorem mem_genWeightSpaceOf (χ : R) (x : L) (m : M) : theorem coe_genWeightSpaceOf_zero (x : L) : ↑(genWeightSpaceOf M (0 : R) x) = ⨆ k, LinearMap.ker (toEnd R L M x ^ k) := by - simp [genWeightSpaceOf, Module.End.maxGenEigenspace_def] + simp [genWeightSpaceOf, ← Module.End.iSup_genEigenspace_eq] /-- If `M` is a representation of a nilpotent Lie algebra `L` and `χ : L → R` is a family of scalars, @@ -267,7 +267,7 @@ lemma genWeightSpaceOf_ne_bot (χ : Weight R L M) (x : L) : lemma hasEigenvalueAt (χ : Weight R L M) (x : L) : (toEnd R L M x).HasEigenvalue (χ x) := by obtain ⟨k : ℕ, hk : (toEnd R L M x).genEigenspace (χ x) k ≠ ⊥⟩ := by - simpa [genWeightSpaceOf, Module.End.maxGenEigenspace_def] using χ.genWeightSpaceOf_ne_bot x + simpa [genWeightSpaceOf, ← Module.End.iSup_genEigenspace_eq] using χ.genWeightSpaceOf_ne_bot x exact Module.End.hasEigenvalue_of_hasGenEigenvalue hk lemma apply_eq_zero_of_isNilpotent [NoZeroSMulDivisors R M] [IsReduced R] @@ -631,8 +631,7 @@ lemma disjoint_genWeightSpaceOf [NoZeroSMulDivisors R M] {x : L} {φ₁ φ₂ : Disjoint (genWeightSpaceOf M φ₁ x) (genWeightSpaceOf M φ₂ x) := by rw [LieSubmodule.disjoint_iff_coe_toSubmodule] dsimp [genWeightSpaceOf] - simp_rw [Module.End.maxGenEigenspace_def] - exact Module.End.disjoint_iSup_genEigenspace _ h + exact Module.End.disjoint_unifEigenspace _ h _ _ lemma disjoint_genWeightSpace [NoZeroSMulDivisors R M] {χ₁ χ₂ : L → R} (h : χ₁ ≠ χ₂) : Disjoint (genWeightSpace M χ₁) (genWeightSpace M χ₂) := by @@ -665,8 +664,7 @@ lemma independent_genWeightSpaceOf [NoZeroSMulDivisors R M] (x : L) : CompleteLattice.Independent fun (χ : R) ↦ genWeightSpaceOf M χ x := by rw [LieSubmodule.independent_iff_coe_toSubmodule] dsimp [genWeightSpaceOf] - simp_rw [Module.End.maxGenEigenspace_def] - exact (toEnd R L M x).independent_genEigenspace + exact (toEnd R L M x).independent_unifEigenspace _ lemma finite_genWeightSpaceOf_ne_bot [NoZeroSMulDivisors R M] [IsNoetherian R M] (x : L) : {χ : R | genWeightSpaceOf M χ x ≠ ⊥}.Finite := @@ -688,16 +686,18 @@ noncomputable instance Weight.instFintype [NoZeroSMulDivisors R M] [IsNoetherian /-- A Lie module `M` of a Lie algebra `L` is triangularizable if the endomorhpism of `M` defined by any `x : L` is triangularizable. -/ class IsTriangularizable : Prop where - iSup_eq_top : ∀ x, ⨆ φ, ⨆ k, (toEnd R L M x).genEigenspace φ k = ⊤ + maxGenEigenspace_eq_top : ∀ x, ⨆ φ, (toEnd R L M x).maxGenEigenspace φ = ⊤ instance (L' : LieSubalgebra R L) [IsTriangularizable R L M] : IsTriangularizable R L' M where - iSup_eq_top x := IsTriangularizable.iSup_eq_top (x : L) + maxGenEigenspace_eq_top x := IsTriangularizable.maxGenEigenspace_eq_top (x : L) instance (I : LieIdeal R L) [IsTriangularizable R L M] : IsTriangularizable R I M where - iSup_eq_top x := IsTriangularizable.iSup_eq_top (x : L) + maxGenEigenspace_eq_top x := IsTriangularizable.maxGenEigenspace_eq_top (x : L) instance [IsTriangularizable R L M] : IsTriangularizable R (LieModule.toEnd R L M).range M where - iSup_eq_top := by rintro ⟨-, x, rfl⟩; exact IsTriangularizable.iSup_eq_top x + maxGenEigenspace_eq_top := by + rintro ⟨-, x, rfl⟩ + exact IsTriangularizable.maxGenEigenspace_eq_top x @[simp] lemma iSup_genWeightSpaceOf_eq_top [IsTriangularizable R L M] (x : L) : @@ -705,8 +705,7 @@ lemma iSup_genWeightSpaceOf_eq_top [IsTriangularizable R L M] (x : L) : rw [← LieSubmodule.coe_toSubmodule_eq_iff, LieSubmodule.iSup_coe_toSubmodule, LieSubmodule.top_coeSubmodule] dsimp [genWeightSpaceOf] - simp_rw [Module.End.maxGenEigenspace_def] - exact IsTriangularizable.iSup_eq_top x + exact IsTriangularizable.maxGenEigenspace_eq_top x open LinearMap Module in @[simp] @@ -729,12 +728,12 @@ variable [Field K] [LieAlgebra K L] [Module K M] [LieModule K L M] [LieAlgebra.I [FiniteDimensional K M] instance instIsTriangularizableOfIsAlgClosed [IsAlgClosed K] : IsTriangularizable K L M := - ⟨fun _ ↦ Module.End.iSup_genEigenspace_eq_top _⟩ + ⟨fun _ ↦ Module.End.iSup_maxGenEigenspace_eq_top _⟩ instance (N : LieSubmodule K L M) [IsTriangularizable K L M] : IsTriangularizable K L N := by refine ⟨fun y ↦ ?_⟩ rw [← N.toEnd_restrict_eq_toEnd y] - exact Module.End.iSup_genEigenspace_restrict_eq_top _ (IsTriangularizable.iSup_eq_top y) + exact Module.End.unifEigenspace_restrict_eq_top _ (IsTriangularizable.maxGenEigenspace_eq_top y) /-- For a triangularizable Lie module in finite dimensions, the weight spaces span the entire space. @@ -745,8 +744,7 @@ lemma iSup_genWeightSpace_eq_top [IsTriangularizable K L M] : LieSubmodule.iInf_coe_toSubmodule, LieSubmodule.top_coeSubmodule, genWeightSpace] 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 + apply IsTriangularizable.maxGenEigenspace_eq_top lemma iSup_genWeightSpace_eq_top' [IsTriangularizable K L M] : ⨆ χ : Weight K L M, genWeightSpace M χ = ⊤ := by diff --git a/Mathlib/Algebra/Lie/Weights/Linear.lean b/Mathlib/Algebra/Lie/Weights/Linear.lean index 741b68fdabb92..2898ac57b5d73 100644 --- a/Mathlib/Algebra/Lie/Weights/Linear.lean +++ b/Mathlib/Algebra/Lie/Weights/Linear.lean @@ -98,17 +98,15 @@ instance instLinearWeightsOfIsLieAbelian [IsLieAbelian L] [NoZeroSMulDivisors R rw [commute_iff_lie_eq, ← LieHom.map_lie, trivial_lie_zero, LieHom.map_zero] intro χ hχ x y simp_rw [Ne, ← LieSubmodule.coe_toSubmodule_eq_iff, genWeightSpace, genWeightSpaceOf, - LieSubmodule.iInf_coe_toSubmodule, LieSubmodule.bot_coeSubmodule, - Module.End.maxGenEigenspace_def] at hχ - exact Module.End.map_add_of_iInf_genEigenspace_ne_bot_of_commute - (toEnd R L M).toLinearMap χ hχ h x y + LieSubmodule.iInf_coe_toSubmodule, LieSubmodule.bot_coeSubmodule] at hχ + exact Module.End.map_add_of_iInf_unifEigenspace_ne_bot_of_commute + (toEnd R L M).toLinearMap χ _ hχ h x y { map_add := aux map_smul := fun χ hχ t x ↦ by simp_rw [Ne, ← LieSubmodule.coe_toSubmodule_eq_iff, genWeightSpace, genWeightSpaceOf, - LieSubmodule.iInf_coe_toSubmodule, LieSubmodule.bot_coeSubmodule, - Module.End.maxGenEigenspace_def] at hχ - exact Module.End.map_smul_of_iInf_genEigenspace_ne_bot - (toEnd R L M).toLinearMap χ hχ t x + LieSubmodule.iInf_coe_toSubmodule, LieSubmodule.bot_coeSubmodule] at hχ + exact Module.End.map_smul_of_iInf_unifEigenspace_ne_bot + (toEnd R L M).toLinearMap χ _ hχ t x map_lie := fun χ hχ t x ↦ by rw [trivial_lie_zero, ← add_left_inj (χ 0), ← aux χ hχ, zero_add, zero_add] } diff --git a/Mathlib/Analysis/InnerProductSpace/JointEigenspace.lean b/Mathlib/Analysis/InnerProductSpace/JointEigenspace.lean index 049ba1532a238..f10d670703e8a 100644 --- a/Mathlib/Analysis/InnerProductSpace/JointEigenspace.lean +++ b/Mathlib/Analysis/InnerProductSpace/JointEigenspace.lean @@ -72,7 +72,7 @@ theorem eigenspace_inf_eigenspace (hAB : A ∘ₗ B = B ∘ₗ A) (γ : 𝕜) : eigenspace A α ⊓ eigenspace B γ = map (Submodule.subtype (eigenspace A α)) (eigenspace (B.restrict (eigenspace_invariant_of_commute hAB α)) γ) := - (eigenspace A α).inf_genEigenspace _ _ (k := 1) + (eigenspace A α).inf_unifEigenspace _ _ (k := 1) variable [FiniteDimensional 𝕜 E] diff --git a/Mathlib/LinearAlgebra/Eigenspace/Basic.lean b/Mathlib/LinearAlgebra/Eigenspace/Basic.lean index d323e2e549474..05a16ab184670 100644 --- a/Mathlib/LinearAlgebra/Eigenspace/Basic.lean +++ b/Mathlib/LinearAlgebra/Eigenspace/Basic.lean @@ -519,10 +519,15 @@ theorem exp_ne_zero_of_hasGenEigenvalue {f : End R M} {μ : R} {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 := by +lemma iSup_genEigenspace_eq (f : End R M) (μ : R) : + ⨆ k, (f.genEigenspace μ) k = f.maxGenEigenspace μ := by simp_rw [maxGenEigenspace, unifEigenspace_top, genEigenspace, OrderHom.coe_mk] +@[deprecated iSup_genEigenspace_eq (since := "2024-10-23")] +lemma maxGenEigenspace_def (f : End R M) (μ : R) : + f.maxGenEigenspace μ = ⨆ k, f.genEigenspace μ k := + (iSup_genEigenspace_eq f μ).symm + theorem genEigenspace_le_maximal (f : End R M) (μ : R) (k : ℕ) : f.genEigenspace μ k ≤ f.maxGenEigenspace μ := (f.unifEigenspace μ).monotone le_top @@ -578,10 +583,12 @@ theorem genEigenspace_le_genEigenspace_finrank [FiniteDimensional K V] (f : End (μ : K) (k : ℕ) : f.genEigenspace μ k ≤ f.genEigenspace μ (finrank K V) := unifEigenspace_le_unifEigenspace_finrank _ _ _ -@[simp] theorem iSup_genEigenspace_eq_genEigenspace_finrank +theorem maxGenEigenspace_eq_genEigenspace_finrank [FiniteDimensional K V] (f : End K V) (μ : K) : - ⨆ k, f.genEigenspace μ k = f.genEigenspace μ (finrank K V) := - le_antisymm (iSup_le (genEigenspace_le_genEigenspace_finrank f μ)) (le_iSup _ _) + f.maxGenEigenspace μ = f.genEigenspace μ (finrank K V) := by + apply le_antisymm _ <| (f.unifEigenspace μ).monotone le_top + rw [unifEigenspace_top_eq_maxUnifEigenspaceIndex] + apply genEigenspace_le_genEigenspace_finrank f μ /-- Generalized eigenspaces for exponents at least `finrank K V` are equal to each other. -/ theorem genEigenspace_eq_genEigenspace_finrank_of_le [FiniteDimensional K V] @@ -590,31 +597,17 @@ theorem genEigenspace_eq_genEigenspace_finrank_of_le [FiniteDimensional K V] 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 - replace h : Commute ((f - μ • (1 : End R M)) ^ k) g := - (h.sub_left <| Algebra.commute_algebraMap_left μ g).pow_left k - intro x hx - simp only [SetLike.mem_coe, mem_genEigenspace] at hx ⊢ - rw [← LinearMap.comp_apply, ← LinearMap.mul_eq_comp, h.eq, LinearMap.mul_eq_comp, - LinearMap.comp_apply, hx, map_zero] - -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 + MapsTo g (f.genEigenspace μ k) (f.genEigenspace μ k) := + mapsTo_unifEigenspace_of_comm h μ k 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⟩ + MapsTo g ↑(f.maxGenEigenspace μ) ↑(f.maxGenEigenspace μ) := + mapsTo_unifEigenspace_of_comm h μ ⊤ +@[deprecated mapsTo_iSup_genEigenspace_of_comm (since := "2024-10-23")] 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] + rw [iSup_genEigenspace_eq] apply mapsTo_maxGenEigenspace_of_comm h /-- The restriction of `f - μ • 1` to the `k`-fold generalized `μ`-eigenspace is nilpotent. -/ @@ -636,7 +629,9 @@ lemma isNilpotent_restrict_maxGenEigenspace_sub_algebraMap [IsNoetherian R M] (f rw [maxGenEigenspace_eq] exact le_rfl +set_option linter.deprecated false in /-- The restriction of `f - μ • 1` to the generalized `μ`-eigenspace is nilpotent. -/ +@[deprecated isNilpotent_restrict_maxGenEigenspace_sub_algebraMap (since := "2024-10-23")] lemma isNilpotent_restrict_iSup_sub_algebraMap [IsNoetherian R M] (f : End R M) (μ : R) (h : MapsTo (f - algebraMap R (End R M) μ) ↑(⨆ k, f.genEigenspace μ k) ↑(⨆ k, f.genEigenspace μ k) := @@ -680,31 +675,40 @@ lemma disjoint_unifEigenspace [NoZeroSMulDivisors R M] isNilpotent_iff_eq_zero, sub_eq_zero] at this contradiction +lemma injOn_unifEigenspace [NoZeroSMulDivisors R M] (f : End R M) (k : ℕ∞) : + InjOn (f.unifEigenspace · k) {μ | f.unifEigenspace μ k ≠ ⊥} := by + rintro μ₁ _ μ₂ hμ₂ hμ₁₂ + by_contra contra + apply hμ₂ + simpa only [hμ₁₂, disjoint_self] using f.disjoint_unifEigenspace contra k k + 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 +@[deprecated disjoint_unifEigenspace (since := "2024-10-23")] lemma disjoint_iSup_genEigenspace [NoZeroSMulDivisors R M] (f : End R M) {μ₁ μ₂ : R} (hμ : μ₁ ≠ μ₂) : Disjoint (⨆ k, f.genEigenspace μ₁ k) (⨆ k, f.genEigenspace μ₂ k) := by simpa only [iSup_genEigenspace_eq] using disjoint_unifEigenspace f hμ ⊤ ⊤ +lemma injOn_maxGenEigenspace [NoZeroSMulDivisors R M] (f : End R M) : + InjOn (f.maxGenEigenspace ·) {μ | f.maxGenEigenspace μ ≠ ⊥} := + injOn_unifEigenspace f ⊤ + +@[deprecated injOn_unifEigenspace (since := "2024-10-23")] lemma injOn_genEigenspace [NoZeroSMulDivisors R M] (f : End R M) : InjOn (⨆ k, f.genEigenspace · k) {μ | ⨆ k, f.genEigenspace μ k ≠ ⊥} := by - rintro μ₁ _ μ₂ hμ₂ (hμ₁₂ : ⨆ k, f.genEigenspace μ₁ k = ⨆ k, f.genEigenspace μ₂ k) - by_contra contra - apply hμ₂ - simpa only [hμ₁₂, disjoint_self] using f.disjoint_iSup_genEigenspace contra + simp_rw [iSup_genEigenspace_eq] + apply injOn_maxGenEigenspace -theorem independent_maxGenEigenspace [NoZeroSMulDivisors R M] (f : End R M) : - CompleteLattice.Independent f.maxGenEigenspace := by +theorem independent_unifEigenspace [NoZeroSMulDivisors R M] (f : End R M) (k : ℕ∞) : + CompleteLattice.Independent (f.unifEigenspace · k) := by classical - suffices ∀ μ (s : Finset R), μ ∉ s → Disjoint (⨆ k, f.genEigenspace μ k) - (s.sup fun μ ↦ ⨆ k, f.genEigenspace μ k) by - show CompleteLattice.Independent (f.maxGenEigenspace ·) - simp_rw [maxGenEigenspace_def, - CompleteLattice.independent_iff_supIndep_of_injOn f.injOn_genEigenspace, + suffices ∀ μ₁ (s : Finset R), μ₁ ∉ s → Disjoint (f.unifEigenspace μ₁ k) + (s.sup fun μ ↦ f.unifEigenspace μ k) by + simp_rw [CompleteLattice.independent_iff_supIndep_of_injOn (injOn_unifEigenspace f k), Finset.supIndep_iff_disjoint_erase] exact fun s μ _ ↦ this _ _ (s.not_mem_erase μ) intro μ₁ s @@ -716,38 +720,43 @@ theorem independent_maxGenEigenspace [NoZeroSMulDivisors R M] (f : End R M) : rw [Finset.sup_insert, disjoint_iff, Submodule.eq_bot_iff] rintro x ⟨hx, hx'⟩ simp only [SetLike.mem_coe] at hx hx' - suffices x ∈ ⨆ k, genEigenspace f μ₂ k by - rw [← Submodule.mem_bot (R := R), ← (f.disjoint_iSup_genEigenspace hμ₁₂).eq_bot] + suffices x ∈ unifEigenspace f μ₂ k by + rw [← Submodule.mem_bot (R := R), ← (f.disjoint_unifEigenspace hμ₁₂ k k).eq_bot] exact ⟨hx, this⟩ obtain ⟨y, hy, z, hz, rfl⟩ := Submodule.mem_sup.mp hx'; clear hx' - let g := f - algebraMap R (End R M) μ₂ - obtain ⟨k : ℕ, hk : (g ^ k) y = 0⟩ := by simpa using hy - have hyz : (g ^ k) (y + z) ∈ - (⨆ k, genEigenspace f μ₁ k) ⊓ s.sup fun μ ↦ ⨆ k, f.genEigenspace μ k := by - refine ⟨f.mapsTo_iSup_genEigenspace_of_comm ?_ μ₁ hx, ?_⟩ - · exact Algebra.mul_sub_algebraMap_pow_commutes f μ₂ k - · rw [SetLike.mem_coe, map_add, hk, zero_add] - suffices (s.sup fun μ ↦ ⨆ k, f.genEigenspace μ k).map (g ^ k) ≤ - s.sup fun μ ↦ ⨆ k, f.genEigenspace μ k by exact this (Submodule.mem_map_of_mem hz) + let g := f - μ₂ • 1 + simp_rw [mem_unifEigenspace, ← exists_prop] at hy ⊢ + peel hy with l hlk hl + simp only [mem_unifEigenspace_nat, LinearMap.mem_ker] at hl + have hyz : (g ^ l) (y + z) ∈ + (f.unifEigenspace μ₁ k) ⊓ s.sup fun μ ↦ f.unifEigenspace μ k := by + refine ⟨f.mapsTo_unifEigenspace_of_comm (g := g ^ l) ?_ μ₁ k hx, ?_⟩ + · exact Algebra.mul_sub_algebraMap_pow_commutes f μ₂ l + · rw [SetLike.mem_coe, map_add, hl, zero_add] + suffices (s.sup fun μ ↦ f.unifEigenspace μ k).map (g ^ l) ≤ + s.sup fun μ ↦ f.unifEigenspace μ k by exact this (Submodule.mem_map_of_mem hz) simp_rw [Finset.sup_eq_iSup, Submodule.map_iSup (ι := R), Submodule.map_iSup (ι := _ ∈ s)] refine iSup₂_mono fun μ _ ↦ ?_ rintro - ⟨u, hu, rfl⟩ - refine f.mapsTo_iSup_genEigenspace_of_comm ?_ μ hu - exact Algebra.mul_sub_algebraMap_pow_commutes f μ₂ k - rw [ih.eq_bot, Submodule.mem_bot] at hyz - simp_rw [Submodule.mem_iSup_of_chain, mem_genEigenspace] - exact ⟨k, hyz⟩ + refine f.mapsTo_unifEigenspace_of_comm ?_ μ k hu + exact Algebra.mul_sub_algebraMap_pow_commutes f μ₂ l + rwa [ih.eq_bot, Submodule.mem_bot] at hyz + +theorem independent_maxGenEigenspace [NoZeroSMulDivisors R M] (f : End R M) : + CompleteLattice.Independent f.maxGenEigenspace := by + apply independent_unifEigenspace +@[deprecated independent_unifEigenspace (since := "2024-10-23")] theorem independent_genEigenspace [NoZeroSMulDivisors R M] (f : End R M) : CompleteLattice.Independent (fun μ ↦ ⨆ k, f.genEigenspace μ k) := by - simp_rw [← maxGenEigenspace_def] + simp_rw [iSup_genEigenspace_eq] 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) : CompleteLattice.Independent f.eigenspace := - f.independent_genEigenspace.mono fun μ ↦ le_iSup (genEigenspace f μ) 1 + (f.independent_unifEigenspace 1).mono fun _ ↦ le_rfl /-- Eigenvectors corresponding to distinct eigenvalues of a linear operator are linearly independent. -/ @@ -781,11 +790,28 @@ theorem genEigenspace_restrict (f : End R M) (p : Submodule R M) (k : ℕ) (μ : · erw [pow_succ, pow_succ, LinearMap.ker_comp, LinearMap.ker_comp, ih, ← LinearMap.ker_comp, LinearMap.comp_assoc] +/-- 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`. -/ +theorem unifEigenspace_restrict (f : End R M) (p : Submodule R M) (k : ℕ∞) (μ : R) + (hfp : ∀ x : M, x ∈ p → f x ∈ p) : + unifEigenspace (LinearMap.restrict f hfp) μ k = + Submodule.comap p.subtype (f.unifEigenspace μ k) := by + ext x + simp only [mem_unifEigenspace, LinearMap.mem_ker, ← mem_genEigenspace, genEigenspace_restrict, + Submodule.mem_comap] + +lemma _root_.Submodule.inf_unifEigenspace (f : End R M) (p : Submodule R M) {k : ℕ∞} {μ : R} + (hfp : ∀ x : M, x ∈ p → f x ∈ p) : + p ⊓ f.unifEigenspace μ k = + (unifEigenspace (LinearMap.restrict f hfp) μ k).map p.subtype := by + rw [f.unifEigenspace_restrict _ _ _ hfp, Submodule.map_comap_eq, Submodule.range_subtype] + +@[deprecated Submodule.inf_unifEigenspace (since := "2024-10-23")] lemma _root_.Submodule.inf_genEigenspace (f : End R M) (p : Submodule R M) {k : ℕ} {μ : R} (hfp : ∀ x : M, x ∈ p → f x ∈ p) : p ⊓ f.genEigenspace μ k = - (genEigenspace (LinearMap.restrict f hfp) μ k).map p.subtype := by - rw [f.genEigenspace_restrict _ _ _ hfp, Submodule.map_comap_eq, Submodule.range_subtype] + (genEigenspace (LinearMap.restrict f hfp) μ k).map p.subtype := + Submodule.inf_unifEigenspace _ _ _ 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} @@ -859,62 +885,97 @@ theorem map_genEigenrange_le {f : End K V} {μ : K} {n : ℕ} : rw [genEigenrange, unifEigenrange_nat] apply LinearMap.map_le_range +lemma unifEigenspace_le_smul (f : Module.End R M) (μ t : R) (k : ℕ∞) : + (f.unifEigenspace μ k) ≤ (t • f).unifEigenspace (t * μ) k := by + intro m hm + simp_rw [mem_unifEigenspace, ← exists_prop, LinearMap.mem_ker] at hm ⊢ + peel hm with l hlk hl + rw [mul_smul, ← smul_sub, smul_pow, LinearMap.smul_apply, hl, smul_zero] + +@[deprecated unifEigenspace_le_smul (since := "2024-10-23")] lemma iSup_genEigenspace_le_smul (f : Module.End R M) (μ t : R) : (⨆ k, f.genEigenspace μ k) ≤ ⨆ k, (t • f).genEigenspace (t * μ) k := by - intro m hm - simp only [Submodule.mem_iSup_of_chain, mem_genEigenspace] at hm ⊢ - refine Exists.imp (fun k hk ↦ ?_) hm - rw [mul_smul, ← smul_sub, smul_pow, LinearMap.smul_apply, hk, smul_zero] + rw [iSup_genEigenspace_eq, iSup_genEigenspace_eq] + apply unifEigenspace_le_smul -lemma iSup_genEigenspace_inf_le_add - (f₁ f₂ : End R M) (μ₁ μ₂ : R) (h : Commute f₁ f₂) : - (⨆ k, f₁.genEigenspace μ₁ k) ⊓ (⨆ k, f₂.genEigenspace μ₂ k) ≤ - ⨆ k, (f₁ + f₂).genEigenspace (μ₁ + μ₂) k := by +lemma unifEigenspace_inf_le_add + (f₁ f₂ : End R M) (μ₁ μ₂ : R) (k₁ k₂ : ℕ∞) (h : Commute f₁ f₂) : + (f₁.unifEigenspace μ₁ k₁) ⊓ (f₂.unifEigenspace μ₂ k₂) ≤ + (f₁ + f₂).unifEigenspace (μ₁ + μ₂) (k₁ + k₂) := by intro m hm - simp only [iSup_le_iff, Submodule.mem_inf, Submodule.mem_iSup_of_chain, - mem_genEigenspace] at hm ⊢ - obtain ⟨⟨k₁, hk₁⟩, ⟨k₂, hk₂⟩⟩ := hm - use k₁ + k₂ - 1 + simp only [Submodule.mem_inf, mem_unifEigenspace, LinearMap.mem_ker] at hm ⊢ + obtain ⟨⟨l₁, hlk₁, hl₁⟩, ⟨l₂, hlk₂, hl₂⟩⟩ := hm + use l₁ + l₂ have : f₁ + f₂ - (μ₁ + μ₂) • 1 = (f₁ - μ₁ • 1) + (f₂ - μ₂ • 1) := by rw [add_smul]; exact add_sub_add_comm f₁ f₂ (μ₁ • 1) (μ₂ • 1) replace h : Commute (f₁ - μ₁ • 1) (f₂ - μ₂ • 1) := (h.sub_right <| Algebra.commute_algebraMap_right μ₂ f₁).sub_left (Algebra.commute_algebraMap_left μ₁ _) rw [this, h.add_pow', LinearMap.coeFn_sum, Finset.sum_apply] + constructor + · simpa only [Nat.cast_add] using add_le_add hlk₁ hlk₂ refine Finset.sum_eq_zero fun ⟨i, j⟩ hij ↦ ?_ suffices (((f₁ - μ₁ • 1) ^ i) * ((f₂ - μ₂ • 1) ^ j)) m = 0 by rw [LinearMap.smul_apply, this, smul_zero] - cases' Nat.le_or_le_of_add_eq_add_pred (Finset.mem_antidiagonal.mp hij) with hi hj - · rw [(h.pow_pow i j).eq, LinearMap.mul_apply, LinearMap.pow_map_zero_of_le hi hk₁, + rw [Finset.mem_antidiagonal] at hij + obtain hi|hj : l₁ ≤ i ∨ l₂ ≤ j := by omega + · rw [(h.pow_pow i j).eq, LinearMap.mul_apply, LinearMap.pow_map_zero_of_le hi hl₁, LinearMap.map_zero] - · rw [LinearMap.mul_apply, LinearMap.pow_map_zero_of_le hj hk₂, LinearMap.map_zero] + · rw [LinearMap.mul_apply, LinearMap.pow_map_zero_of_le hj hl₂, LinearMap.map_zero] -lemma map_smul_of_iInf_genEigenspace_ne_bot [NoZeroSMulDivisors R M] +@[deprecated unifEigenspace_inf_le_add (since := "2024-10-23")] +lemma iSup_genEigenspace_inf_le_add + (f₁ f₂ : End R M) (μ₁ μ₂ : R) (h : Commute f₁ f₂) : + (⨆ k, f₁.genEigenspace μ₁ k) ⊓ (⨆ k, f₂.genEigenspace μ₂ k) ≤ + ⨆ k, (f₁ + f₂).genEigenspace (μ₁ + μ₂) k := by + simp_rw [iSup_genEigenspace_eq] + apply unifEigenspace_inf_le_add + assumption + +lemma map_smul_of_iInf_unifEigenspace_ne_bot [NoZeroSMulDivisors R M] {L F : Type*} [SMul R L] [FunLike F L (End R M)] [MulActionHomClass F R L (End R M)] (f : F) - (μ : L → R) (h_ne : ⨅ x, ⨆ k, (f x).genEigenspace (μ x) k ≠ ⊥) + (μ : L → R) (k : ℕ∞) (h_ne : ⨅ x, (f x).unifEigenspace (μ x) k ≠ ⊥) (t : R) (x : L) : μ (t • x) = t • μ x := by by_contra contra - let g : L → Submodule R M := fun x ↦ ⨆ k, (f x).genEigenspace (μ x) k + let g : L → Submodule R M := fun x ↦ (f x).unifEigenspace (μ x) k have : ⨅ x, g x ≤ g x ⊓ g (t • x) := le_inf_iff.mpr ⟨iInf_le g x, iInf_le g (t • x)⟩ refine h_ne <| eq_bot_iff.mpr (le_trans this (disjoint_iff_inf_le.mp ?_)) - apply Disjoint.mono_left (iSup_genEigenspace_le_smul (f x) (μ x) t) + apply Disjoint.mono_left (unifEigenspace_le_smul (f x) (μ x) t k) simp only [g, map_smul] - exact disjoint_iSup_genEigenspace (t • f x) (Ne.symm contra) + exact disjoint_unifEigenspace (t • f x) (Ne.symm contra) k k -lemma map_add_of_iInf_genEigenspace_ne_bot_of_commute [NoZeroSMulDivisors R M] - {L F : Type*} [Add L] [FunLike F L (End R M)] [AddHomClass F L (End R M)] (f : F) +@[deprecated map_smul_of_iInf_unifEigenspace_ne_bot (since := "2024-10-23")] +lemma map_smul_of_iInf_genEigenspace_ne_bot [NoZeroSMulDivisors R M] + {L F : Type*} [SMul R L] [FunLike F L (End R M)] [MulActionHomClass F R L (End R M)] (f : F) (μ : L → R) (h_ne : ⨅ x, ⨆ k, (f x).genEigenspace (μ x) k ≠ ⊥) + (t : R) (x : L) : + μ (t • x) = t • μ x := by + simp_rw [iSup_genEigenspace_eq] at h_ne + apply map_smul_of_iInf_unifEigenspace_ne_bot f μ ⊤ h_ne t x + +lemma map_add_of_iInf_unifEigenspace_ne_bot_of_commute [NoZeroSMulDivisors R M] + {L F : Type*} [Add L] [FunLike F L (End R M)] [AddHomClass F L (End R M)] (f : F) + (μ : L → R) (k : ℕ∞) (h_ne : ⨅ x, (f x).unifEigenspace (μ x) k ≠ ⊥) (h : ∀ x y, Commute (f x) (f y)) (x y : L) : μ (x + y) = μ x + μ y := by by_contra contra - let g : L → Submodule R M := fun x ↦ ⨆ k, (f x).genEigenspace (μ x) k + let g : L → Submodule R M := fun x ↦ (f x).unifEigenspace (μ x) k have : ⨅ x, g x ≤ (g x ⊓ g y) ⊓ g (x + y) := le_inf_iff.mpr ⟨le_inf_iff.mpr ⟨iInf_le g x, iInf_le g y⟩, iInf_le g (x + y)⟩ refine h_ne <| eq_bot_iff.mpr (le_trans this (disjoint_iff_inf_le.mp ?_)) - apply Disjoint.mono_left (iSup_genEigenspace_inf_le_add (f x) (f y) (μ x) (μ y) (h x y)) + apply Disjoint.mono_left (unifEigenspace_inf_le_add (f x) (f y) (μ x) (μ y) k k (h x y)) simp only [g, map_add] - exact disjoint_iSup_genEigenspace (f x + f y) (Ne.symm contra) + exact disjoint_unifEigenspace (f x + f y) (Ne.symm contra) _ k + +@[deprecated map_add_of_iInf_unifEigenspace_ne_bot_of_commute (since := "2024-10-23")] +lemma map_add_of_iInf_genEigenspace_ne_bot_of_commute [NoZeroSMulDivisors R M] + {L F : Type*} [Add L] [FunLike F L (End R M)] [AddHomClass F L (End R M)] (f : F) + (μ : L → R) (h_ne : ⨅ x, ⨆ k, (f x).genEigenspace (μ x) k ≠ ⊥) + (h : ∀ x y, Commute (f x) (f y)) (x y : L) : + μ (x + y) = μ x + μ y := by + simp_rw [iSup_genEigenspace_eq] at h_ne + apply map_add_of_iInf_unifEigenspace_ne_bot_of_commute f μ ⊤ h_ne h x y end End diff --git a/Mathlib/LinearAlgebra/Eigenspace/Pi.lean b/Mathlib/LinearAlgebra/Eigenspace/Pi.lean index 3eeacc0153eb7..2980b51c48875 100644 --- a/Mathlib/LinearAlgebra/Eigenspace/Pi.lean +++ b/Mathlib/LinearAlgebra/Eigenspace/Pi.lean @@ -45,8 +45,7 @@ lemma _root_.Submodule.inf_iInf_maxGenEigenspace_of_forall_mapsTo {μ : ι → R (⨅ 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_def, ((f _).genEigenspace _).mono.directed_le.inf_iSup_eq, - p.inf_genEigenspace _ (hfp _), ← Submodule.map_iSup, Submodule.map_iInf _ p.injective_subtype] + · simp_rw [inf_iInf, p.inf_unifEigenspace _ (hfp _), Submodule.map_iInf _ p.injective_subtype] /-- Given a family of endomorphisms `i ↦ f i`, a family of candidate eigenvalues `i ↦ μ i`, and a distinguished index `i` whose maximal generalised `μ i`-eigenspace is invariant wrt every `f j`, @@ -64,8 +63,10 @@ lemma iInf_maxGenEigenspace_restrict_map_subtype_eq refine le_antisymm ?_ inf_le_right simpa only [le_inf_iff, le_refl, and_true] using iInf_le _ _ 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 _)] + conv_rhs => + enter [1] + ext + rw [p.inf_unifEigenspace (f _) (h _)] variable [NoZeroSMulDivisors R M] @@ -166,8 +167,7 @@ lemma iSup_iInf_maxGenEigenspace_eq_top_of_forall_mapsTo [FiniteDimensional K M] apply ih _ (hy φ) · intro j k μ exact mapsTo_restrict_maxGenEigenspace_restrict_of_mapsTo (f j) (f k) _ _ (h j k μ) - · simp_rw [maxGenEigenspace_def] at h' ⊢ - exact fun j ↦ Module.End.iSup_genEigenspace_restrict_eq_top _ (h' j) + · exact fun j ↦ Module.End.unifEigenspace_restrict_eq_top _ (h' j) · rfl replace ih (φ : K) : ⨆ (χ : ι → K) (_ : χ i = φ), ⨅ j, maxGenEigenspace ((f j).restrict (hi j φ)) (χ j) = ⊤ := by @@ -177,8 +177,7 @@ lemma iSup_iInf_maxGenEigenspace_eq_top_of_forall_mapsTo [FiniteDimensional K M] 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) + exact ((f i).disjoint_unifEigenspace hχ.symm _ _).mono_right (iInf_le _ i) replace ih (φ : K) : ⨆ (χ : ι → K) (_ : χ i = φ), ⨅ j, maxGenEigenspace (f j) (χ j) = maxGenEigenspace (f i) φ := by diff --git a/Mathlib/LinearAlgebra/Eigenspace/Semisimple.lean b/Mathlib/LinearAlgebra/Eigenspace/Semisimple.lean index 96229aa54786f..26ccc18c96849 100644 --- a/Mathlib/LinearAlgebra/Eigenspace/Semisimple.lean +++ b/Mathlib/LinearAlgebra/Eigenspace/Semisimple.lean @@ -26,17 +26,20 @@ namespace Module.End variable {R M : Type*} [CommRing R] [AddCommGroup M] [Module R M] {f g : End R M} lemma apply_eq_of_mem_of_comm_of_isFinitelySemisimple_of_isNil - {μ : R} {k : ℕ} {m : M} (hm : m ∈ f.genEigenspace μ k) + {μ : R} {k : ℕ∞} {m : M} (hm : m ∈ f.unifEigenspace μ k) (hfg : Commute f g) (hss : g.IsFinitelySemisimple) (hnil : IsNilpotent (f - g)) : g m = μ • m := by - set p := f.genEigenspace μ k - have h₁ : MapsTo g p p := mapsTo_genEigenspace_of_comm hfg μ k + rw [f.mem_unifEigenspace] at hm + obtain ⟨l, -, hm⟩ := hm + rw [LinearMap.mem_ker, ← f.mem_genEigenspace] at hm + set p := f.genEigenspace μ l + have h₁ : MapsTo g p p := mapsTo_unifEigenspace_of_comm hfg μ l have h₂ : MapsTo (g - algebraMap R (End R M) μ) p p := - mapsTo_genEigenspace_of_comm (hfg.sub_right <| Algebra.commute_algebraMap_right μ f) μ k + mapsTo_unifEigenspace_of_comm (hfg.sub_right <| Algebra.commute_algebraMap_right μ f) μ l have h₃ : MapsTo (f - g) p p := - mapsTo_genEigenspace_of_comm (Commute.sub_right rfl hfg) μ k + mapsTo_unifEigenspace_of_comm (Commute.sub_right rfl hfg) μ l have h₄ : MapsTo (f - algebraMap R (End R M) μ) p p := - mapsTo_genEigenspace_of_comm (Algebra.mul_sub_algebraMap_commutes f μ) μ k + mapsTo_unifEigenspace_of_comm (Algebra.mul_sub_algebraMap_commutes f μ) μ l replace hfg : Commute (f - algebraMap R (End R M) μ) (f - g) := (Commute.sub_right rfl hfg).sub_left <| Algebra.commute_algebraMap_left μ (f - g) suffices IsNilpotent ((g - algebraMap R (End R M) μ).restrict h₂) by @@ -44,7 +47,15 @@ lemma apply_eq_of_mem_of_comm_of_isFinitelySemisimple_of_isNil eq_zero_of_isNilpotent_of_isFinitelySemisimple this (by simpa using hss.restrict _) simpa [LinearMap.restrict_apply, sub_eq_zero] using LinearMap.congr_fun this ⟨m, hm⟩ simpa [LinearMap.restrict_sub h₄ h₃] using (LinearMap.restrict_commute hfg h₄ h₃).isNilpotent_sub - (f.isNilpotent_restrict_sub_algebraMap μ k) (Module.End.isNilpotent.restrict h₃ hnil) + (f.isNilpotent_restrict_sub_algebraMap μ l) (Module.End.isNilpotent.restrict h₃ hnil) + +lemma IsFinitelySemisimple.unifEigenspace_eq_eigenspace + (hf : f.IsFinitelySemisimple) (μ : R) {k : ℕ∞} (hk : 0 < k) : + f.unifEigenspace μ k = f.eigenspace μ := by + refine le_antisymm (fun m hm ↦ mem_eigenspace_iff.mpr ?_) (f.unifEigenspace μ |>.mono ?_) + · apply apply_eq_of_mem_of_comm_of_isFinitelySemisimple_of_isNil hm rfl hf + simp + · exact Order.one_le_iff_pos.mpr hk lemma IsFinitelySemisimple.genEigenspace_eq_eigenspace (hf : f.IsFinitelySemisimple) (μ : R) {k : ℕ} (hk : 0 < k) : @@ -54,8 +65,7 @@ lemma IsFinitelySemisimple.genEigenspace_eq_eigenspace lemma IsFinitelySemisimple.maxGenEigenspace_eq_eigenspace (hf : f.IsFinitelySemisimple) (μ : R) : - f.maxGenEigenspace μ = f.eigenspace μ := by - simp_rw [maxGenEigenspace_def, ← (f.genEigenspace μ).monotone.iSup_nat_add 1, - hf.genEigenspace_eq_eigenspace μ (Nat.zero_lt_succ _), ciSup_const] + f.maxGenEigenspace μ = f.eigenspace μ := + hf.unifEigenspace_eq_eigenspace μ ENat.top_pos end Module.End diff --git a/Mathlib/LinearAlgebra/Eigenspace/Triangularizable.lean b/Mathlib/LinearAlgebra/Eigenspace/Triangularizable.lean index e619e16819564..4ac57d9a9c04a 100644 --- a/Mathlib/LinearAlgebra/Eigenspace/Triangularizable.lean +++ b/Mathlib/LinearAlgebra/Eigenspace/Triangularizable.lean @@ -45,16 +45,20 @@ variable {K V : Type*} [Field K] [AddCommGroup V] [Module K V] namespace Module.End +theorem exists_hasEigenvalue_of_unifEigenspace_eq_top [Nontrivial M] {f : End R M} (k : ℕ∞) + (hf : ⨆ μ, f.unifEigenspace μ k = ⊤) : + ∃ μ, f.HasEigenvalue μ := by + suffices ∃ μ, f.HasUnifEigenvalue μ k by + peel this with μ hμ + exact HasUnifEigenvalue.lt zero_lt_one hμ + simp [HasUnifEigenvalue, ← not_forall, ← iSup_eq_bot, hf] + +@[deprecated exists_hasEigenvalue_of_unifEigenspace_eq_top (since := "2024-10-11")] theorem exists_hasEigenvalue_of_iSup_genEigenspace_eq_top [Nontrivial M] {f : End R M} (hf : ⨆ μ, ⨆ k, f.genEigenspace μ k = ⊤) : ∃ μ, f.HasEigenvalue μ := by - by_contra! contra - suffices ∀ μ, ⨆ k, f.genEigenspace μ k = ⊥ by simp [this] at hf - intro μ - replace contra : ∀ k, f.genEigenspace μ k = ⊥ := fun k ↦ by - have hk : ¬ f.HasGenEigenvalue μ k := fun hk ↦ contra μ (f.hasEigenvalue_of_hasGenEigenvalue hk) - rwa [hasGenEigenvalue_iff, not_not] at hk - simp [contra] + simp_rw [iSup_genEigenspace_eq] at hf + apply exists_hasEigenvalue_of_unifEigenspace_eq_top _ hf -- This is Lemma 5.21 of [axler2015], although we are no longer following that proof. /-- In finite dimensions, over an algebraically closed field, every linear endomorphism has an @@ -71,8 +75,8 @@ noncomputable instance [IsAlgClosed K] [FiniteDimensional K V] [Nontrivial V] (f -- Lemma 8.21 of [axler2015] /-- In finite dimensions, over an algebraically closed field, the generalized eigenspaces of any linear endomorphism span the whole space. -/ -theorem iSup_genEigenspace_eq_top [IsAlgClosed K] [FiniteDimensional K V] (f : End K V) : - ⨆ (μ : K) (k : ℕ), f.genEigenspace μ k = ⊤ := by +theorem iSup_maxGenEigenspace_eq_top [IsAlgClosed K] [FiniteDimensional K V] (f : End K V) : + ⨆ (μ : K), f.maxGenEigenspace μ = ⊤ := by -- We prove the claim by strong induction on the dimension of the vector space. induction' h_dim : finrank K V using Nat.strong_induction_on with n ih generalizing V cases' n with n @@ -105,57 +109,68 @@ theorem iSup_genEigenspace_eq_top [IsAlgClosed K] [FiniteDimensional K V] (f : E -- Therefore the dimension `ER` mus be smaller than `finrank K V`. have h_dim_ER : finrank K ER < n.succ := by linarith -- This allows us to apply the induction hypothesis on `ER`: - have ih_ER : ⨆ (μ : K) (k : ℕ), f'.genEigenspace μ k = ⊤ := + have ih_ER : ⨆ (μ : K), f'.maxGenEigenspace μ = ⊤ := ih (finrank K ER) h_dim_ER f' rfl -- The induction hypothesis gives us a statement about subspaces of `ER`. We can transfer this -- to a statement about subspaces of `V` via `Submodule.subtype`: - have ih_ER' : ⨆ (μ : K) (k : ℕ), (f'.genEigenspace μ k).map ER.subtype = ER := by + have ih_ER' : ⨆ (μ : K), (f'.maxGenEigenspace μ).map ER.subtype = ER := by simp only [(Submodule.map_iSup _ _).symm, ih_ER, Submodule.map_subtype_top ER] -- Moreover, every generalized eigenspace of `f'` is contained in the corresponding generalized -- eigenspace of `f`. have hff' : - ∀ μ k, (f'.genEigenspace μ k).map ER.subtype ≤ f.genEigenspace μ k := by + ∀ μ, (f'.maxGenEigenspace μ).map ER.subtype ≤ f.maxGenEigenspace μ := by intros - rw [genEigenspace_restrict] + rw [maxGenEigenspace, unifEigenspace_restrict] apply Submodule.map_comap_le -- It follows that `ER` is contained in the span of all generalized eigenvectors. - have hER : ER ≤ ⨆ (μ : K) (k : ℕ), f.genEigenspace μ k := by + have hER : ER ≤ ⨆ (μ : K), f.maxGenEigenspace μ := by rw [← ih_ER'] - exact iSup₂_mono hff' + exact iSup_mono hff' -- `ES` is contained in this span by definition. - have hES : ES ≤ ⨆ (μ : K) (k : ℕ), f.genEigenspace μ k := - le_trans (le_iSup (fun k => f.genEigenspace μ₀ k) (finrank K V)) - (le_iSup (fun μ : K => ⨆ k : ℕ, f.genEigenspace μ k) μ₀) + have hES : ES ≤ ⨆ (μ : K), f.maxGenEigenspace μ := + ((f.unifEigenspace μ₀).mono le_top).trans (le_iSup f.maxGenEigenspace μ₀) -- Moreover, we know that `ER` and `ES` are disjoint. have h_disjoint : Disjoint ER ES := generalized_eigenvec_disjoint_range_ker f μ₀ -- Since the dimensions of `ER` and `ES` add up to the dimension of `V`, it follows that the -- span of all generalized eigenvectors is all of `V`. - show ⨆ (μ : K) (k : ℕ), f.genEigenspace μ k = ⊤ + show ⨆ (μ : K), f.maxGenEigenspace μ = ⊤ rw [← top_le_iff, ← Submodule.eq_top_of_disjoint ER ES h_dim_add h_disjoint] apply sup_le hER hES +-- Lemma 8.21 of [axler2015] +/-- In finite dimensions, over an algebraically closed field, the generalized eigenspaces of any +linear endomorphism span the whole space. -/ +@[deprecated iSup_maxGenEigenspace_eq_top (since := "2024-10-11")] +theorem iSup_genEigenspace_eq_top [IsAlgClosed K] [FiniteDimensional K V] (f : End K V) : + ⨆ (μ : K) (k : ℕ), f.genEigenspace μ k = ⊤ := by + simp_rw [iSup_genEigenspace_eq] + apply iSup_maxGenEigenspace_eq_top + end Module.End namespace Submodule variable {p : Submodule K V} {f : Module.End K V} -theorem inf_iSup_genEigenspace [FiniteDimensional K V] (h : ∀ x ∈ p, f x ∈ p) : - p ⊓ ⨆ μ, ⨆ k, f.genEigenspace μ k = ⨆ μ, ⨆ k, p ⊓ f.genEigenspace μ k := by - simp_rw [← (f.genEigenspace _).mono.directed_le.inf_iSup_eq] +theorem inf_iSup_unifEigenspace [FiniteDimensional K V] (h : ∀ x ∈ p, f x ∈ p) (k : ℕ∞) : + p ⊓ ⨆ μ, f.unifEigenspace μ k = ⨆ μ, p ⊓ f.unifEigenspace μ k := by refine le_antisymm (fun m hm ↦ ?_) (le_inf_iff.mpr ⟨iSup_le fun μ ↦ inf_le_left, iSup_mono fun μ ↦ inf_le_right⟩) classical - obtain ⟨hm₀ : m ∈ p, hm₁ : m ∈ ⨆ μ, ⨆ k, f.genEigenspace μ k⟩ := hm + obtain ⟨hm₀ : m ∈ p, hm₁ : m ∈ ⨆ μ, f.unifEigenspace μ k⟩ := hm obtain ⟨m, hm₂, rfl⟩ := (mem_iSup_iff_exists_finsupp _ _).mp hm₁ suffices ∀ μ, (m μ : V) ∈ p by exact (mem_iSup_iff_exists_finsupp _ _).mpr ⟨m, fun μ ↦ mem_inf.mp ⟨this μ, hm₂ μ⟩, rfl⟩ intro μ by_cases hμ : μ ∈ m.support; swap · simp only [Finsupp.not_mem_support_iff.mp hμ, p.zero_mem] + have hm₂_aux := hm₂ + simp_rw [Module.End.mem_unifEigenspace] at hm₂_aux + choose l hlk hl using hm₂_aux + let l₀ : ℕ := m.support.sup l have h_comm : ∀ (μ₁ μ₂ : K), - Commute ((f - algebraMap K (End K V) μ₁) ^ finrank K V) - ((f - algebraMap K (End K V) μ₂) ^ finrank K V) := fun μ₁ μ₂ ↦ + Commute ((f - algebraMap K (End K V) μ₁) ^ l₀) + ((f - algebraMap K (End K V) μ₂) ^ l₀) := fun μ₁ μ₂ ↦ ((Commute.sub_right rfl <| Algebra.commute_algebraMap_right _ _).sub_left (Algebra.commute_algebraMap_left _ _)).pow_pow _ _ let g : End K V := (m.support.erase μ).noncommProd _ fun μ₁ _ μ₂ _ _ ↦ h_comm μ₁ μ₂ @@ -168,42 +183,72 @@ theorem inf_iSup_genEigenspace [FiniteDimensional K V] (h : ∀ x ∈ p, f x ∈ rintro μ' hμ' split_ifs with hμμ' · rw [hμμ'] - replace hm₂ : ((f - algebraMap K (End K V) μ') ^ finrank K V) (m μ') = 0 := by - obtain ⟨k, hk⟩ := (mem_iSup_of_chain _ _).mp (hm₂ μ') - simpa only [End.mem_genEigenspace] using - Module.End.genEigenspace_le_genEigenspace_finrank _ _ k hk + have hl₀ : ((f - algebraMap K (End K V) μ') ^ l₀) (m μ') = 0 := by + rw [← LinearMap.mem_ker, Algebra.algebraMap_eq_smul_one, ← End.mem_unifEigenspace_nat] + simp_rw [← End.mem_unifEigenspace_nat] at hl + suffices (l μ' : ℕ∞) ≤ l₀ from (f.unifEigenspace μ').mono this (hl μ') + simpa only [Nat.cast_le] using Finset.le_sup hμ' have : _ = g := (m.support.erase μ).noncommProd_erase_mul (Finset.mem_erase.mpr ⟨hμμ', hμ'⟩) - (fun μ ↦ (f - algebraMap K (End K V) μ) ^ finrank K V) (fun μ₁ _ μ₂ _ _ ↦ h_comm μ₁ μ₂) - rw [← this, LinearMap.mul_apply, hm₂, _root_.map_zero] + (fun μ ↦ (f - algebraMap K (End K V) μ) ^ l₀) (fun μ₁ _ μ₂ _ _ ↦ h_comm μ₁ μ₂) + rw [← this, LinearMap.mul_apply, hl₀, _root_.map_zero] have hg₁ : MapsTo g p p := Finset.noncommProd_induction _ _ _ (fun g' : End K V ↦ MapsTo g' p p) (fun f₁ f₂ ↦ MapsTo.comp) (mapsTo_id _) fun μ' _ ↦ by suffices MapsTo (f - algebraMap K (End K V) μ') p p by - simp only [LinearMap.coe_pow]; exact this.iterate (finrank K V) + simp only [LinearMap.coe_pow]; exact this.iterate l₀ intro x hx rw [LinearMap.sub_apply, algebraMap_end_apply] exact p.sub_mem (h _ hx) (smul_mem p μ' hx) - have hg₂ : MapsTo g ↑(⨆ k, f.genEigenspace μ k) ↑(⨆ k, f.genEigenspace μ k) := - f.mapsTo_iSup_genEigenspace_of_comm hfg μ - have hg₃ : InjOn g ↑(⨆ k, f.genEigenspace μ k) := by + have hg₂ : MapsTo g ↑(f.unifEigenspace μ k) ↑(f.unifEigenspace μ k) := + f.mapsTo_unifEigenspace_of_comm hfg μ k + have hg₃ : InjOn g ↑(f.unifEigenspace μ k) := by apply LinearMap.injOn_of_disjoint_ker (subset_refl _) - have this := f.independent_genEigenspace - simp_rw [f.iSup_genEigenspace_eq_genEigenspace_finrank] at this ⊢ + have this := f.independent_unifEigenspace k + have aux (μ') (_hμ' : μ' ∈ m.support.erase μ) : + (f.unifEigenspace μ') ↑l₀ ≤ (f.unifEigenspace μ') k := by + apply (f.unifEigenspace μ').mono + rintro k rfl + simp only [ENat.some_eq_coe, Nat.cast_inj, exists_eq_left'] + apply Finset.sup_le + intro i _hi + simpa using hlk i rw [LinearMap.ker_noncommProd_eq_of_supIndep_ker, ← Finset.sup_eq_iSup] - · simpa only [End.genEigenspace_def] using - Finset.supIndep_iff_disjoint_erase.mp (this.supIndep' m.support) μ hμ - · simpa only [End.genEigenspace_def] using this.supIndep' (m.support.erase μ) + · have := Finset.supIndep_iff_disjoint_erase.mp (this.supIndep' m.support) μ hμ + apply this.mono_right + apply Finset.sup_mono_fun + intro μ' hμ' + rw [Algebra.algebraMap_eq_smul_one, ← End.unifEigenspace_nat] + apply aux μ' hμ' + · have := this.supIndep' (m.support.erase μ) + apply Finset.supIndep_antimono_fun _ this + intro μ' hμ' + rw [Algebra.algebraMap_eq_smul_one, ← End.unifEigenspace_nat] + apply aux μ' hμ' have hg₄ : SurjOn g - ↑(p ⊓ ⨆ k, f.genEigenspace μ k) ↑(p ⊓ ⨆ k, f.genEigenspace μ k) := by + ↑(p ⊓ f.unifEigenspace μ k) ↑(p ⊓ f.unifEigenspace μ k) := by have : MapsTo g - ↑(p ⊓ ⨆ k, f.genEigenspace μ k) ↑(p ⊓ ⨆ k, f.genEigenspace μ k) := + ↑(p ⊓ f.unifEigenspace μ k) ↑(p ⊓ f.unifEigenspace μ k) := hg₁.inter_inter hg₂ rw [← LinearMap.injOn_iff_surjOn this] exact hg₃.mono inter_subset_right specialize hm₂ μ - obtain ⟨y, ⟨hy₀ : y ∈ p, hy₁ : y ∈ ⨆ k, f.genEigenspace μ k⟩, hy₂ : g y = g (m μ)⟩ := + obtain ⟨y, ⟨hy₀ : y ∈ p, hy₁ : y ∈ f.unifEigenspace μ k⟩, hy₂ : g y = g (m μ)⟩ := hg₄ ⟨(hg₀ ▸ hg₁ hm₀), hg₂ hm₂⟩ rwa [← hg₃ hy₁ hm₂ hy₂] +set_option linter.deprecated false in +@[deprecated inf_iSup_unifEigenspace (since := "2024-10-11")] +theorem inf_iSup_genEigenspace [FiniteDimensional K V] (h : ∀ x ∈ p, f x ∈ p) : + p ⊓ ⨆ μ, ⨆ k, f.genEigenspace μ k = ⨆ μ, ⨆ k, p ⊓ f.genEigenspace μ k := by + simp_rw [← (f.genEigenspace _).mono.directed_le.inf_iSup_eq, f.iSup_genEigenspace_eq] + apply inf_iSup_unifEigenspace h ⊤ + +theorem eq_iSup_inf_unifEigenspace [FiniteDimensional K V] (k : ℕ∞) + (h : ∀ x ∈ p, f x ∈ p) (h' : ⨆ μ, f.unifEigenspace μ k = ⊤) : + p = ⨆ μ, p ⊓ f.unifEigenspace μ k := by + rw [← inf_iSup_unifEigenspace h, h', inf_top_eq] + +set_option linter.deprecated false in +@[deprecated eq_iSup_inf_unifEigenspace (since := "2024-10-11")] theorem eq_iSup_inf_genEigenspace [FiniteDimensional K V] (h : ∀ x ∈ p, f x ∈ p) (h' : ⨆ μ, ⨆ k, f.genEigenspace μ k = ⊤) : p = ⨆ μ, ⨆ k, p ⊓ f.genEigenspace μ k := by @@ -213,12 +258,22 @@ end Submodule /-- In finite dimensions, if the generalized eigenspaces of a linear endomorphism span the whole space then the same is true of its restriction to any invariant submodule. -/ +theorem Module.End.unifEigenspace_restrict_eq_top + {p : Submodule K V} {f : Module.End K V} [FiniteDimensional K V] {k : ℕ∞} + (h : ∀ x ∈ p, f x ∈ p) (h' : ⨆ μ, f.unifEigenspace μ k = ⊤) : + ⨆ μ, Module.End.unifEigenspace (LinearMap.restrict f h) μ k = ⊤ := by + have := congr_arg (Submodule.comap p.subtype) (Submodule.eq_iSup_inf_unifEigenspace k h h') + have h_inj : Function.Injective p.subtype := Subtype.coe_injective + simp_rw [Submodule.inf_unifEigenspace f p h, Submodule.comap_subtype_self, + ← Submodule.map_iSup, Submodule.comap_map_eq_of_injective h_inj] at this + exact this.symm + +/-- In finite dimensions, if the generalized eigenspaces of a linear endomorphism span the whole +space then the same is true of its restriction to any invariant submodule. -/ +@[deprecated Module.End.unifEigenspace_restrict_eq_top (since := "2024-10-11")] theorem Module.End.iSup_genEigenspace_restrict_eq_top {p : Submodule K V} {f : Module.End K V} [FiniteDimensional K V] (h : ∀ x ∈ p, f x ∈ p) (h' : ⨆ μ, ⨆ k, f.genEigenspace μ k = ⊤) : ⨆ μ, ⨆ k, Module.End.genEigenspace (LinearMap.restrict f h) μ k = ⊤ := by - have := congr_arg (Submodule.comap p.subtype) (Submodule.eq_iSup_inf_genEigenspace h h') - have h_inj : Function.Injective p.subtype := Subtype.coe_injective - simp_rw [Submodule.inf_genEigenspace f p h, Submodule.comap_subtype_self, - ← Submodule.map_iSup, Submodule.comap_map_eq_of_injective h_inj] at this - exact this.symm + simp_rw [iSup_genEigenspace_eq] at h' ⊢ + apply Module.End.unifEigenspace_restrict_eq_top h h' diff --git a/Mathlib/LinearAlgebra/Eigenspace/Zero.lean b/Mathlib/LinearAlgebra/Eigenspace/Zero.lean index 30837353e1b0a..cbae5ab149388 100644 --- a/Mathlib/LinearAlgebra/Eigenspace/Zero.lean +++ b/Mathlib/LinearAlgebra/Eigenspace/Zero.lean @@ -132,7 +132,7 @@ lemma finrank_maxGenEigenspace (φ : Module.End K M) : finrank K (φ.maxGenEigenspace 0) = natTrailingDegree (φ.charpoly) := by set V := φ.maxGenEigenspace 0 have hV : V = ⨆ (n : ℕ), ker (φ ^ n) := by - simp [V, Module.End.maxGenEigenspace_def, Module.End.genEigenspace_def] + simp [V, ← Module.End.iSup_genEigenspace_eq, Module.End.genEigenspace_def] let W := ⨅ (n : ℕ), LinearMap.range (φ ^ n) have hVW : IsCompl V W := by rw [hV] diff --git a/Mathlib/Order/SupIndep.lean b/Mathlib/Order/SupIndep.lean index cb3c3c12cac6a..19f1467d5960c 100644 --- a/Mathlib/Order/SupIndep.lean +++ b/Mathlib/Order/SupIndep.lean @@ -94,6 +94,29 @@ theorem supIndep_iff_disjoint_erase [DecidableEq ι] : ⟨fun hs _ hi => hs (erase_subset _ _) hi (not_mem_erase _ _), fun hs _ ht i hi hit => (hs i hi).mono_right (sup_mono fun _ hj => mem_erase.2 ⟨ne_of_mem_of_not_mem hj hit, ht hj⟩)⟩ +theorem supIndep_antimono_fun {g : ι → α} (h : ∀ x ∈ s, f x ≤ g x) (h : s.SupIndep g) : + s.SupIndep f := by + classical + induction s using Finset.induction_on with + | empty => apply Finset.supIndep_empty + | @insert i s his IH => + rename_i hle + rw [Finset.supIndep_iff_disjoint_erase] at h ⊢ + intro j hj + simp_all only [Finset.mem_insert, or_true, implies_true, true_implies, forall_eq_or_imp, + Finset.erase_insert_eq_erase, not_false_eq_true, Finset.erase_eq_of_not_mem] + obtain rfl | hj := hj + · simp only [Finset.erase_insert_eq_erase] + apply h.left.mono hle.left + apply (Finset.sup_mono _).trans (Finset.sup_mono_fun hle.right) + exact Finset.erase_subset _ _ + · apply (h.right j hj).mono (hle.right j hj) (Finset.sup_mono_fun _) + intro k hk + simp only [Finset.mem_erase, ne_eq, Finset.mem_insert] at hk + obtain ⟨-, rfl | hk⟩ := hk + · exact hle.left + · exact hle.right k hk + theorem SupIndep.image [DecidableEq ι] {s : Finset ι'} {g : ι' → ι} (hs : s.SupIndep (f ∘ g)) : (s.image g).SupIndep f := by intro t ht i hi hit From cf1c5c28f832f1c486a60c424d5c17823897d778 Mon Sep 17 00:00:00 2001 From: sven-manthe Date: Fri, 25 Oct 2024 20:58:33 +0000 Subject: [PATCH 429/505] chore: add refl annotations to List.IsPrefix etc (#18221) These lemmas are currently not working with the rfl tactic even after import Mathlib. Since they are declared in the core, it seems easiest to just provide the annotation afterwards. --- Mathlib/Data/List/Infix.lean | 1 + 1 file changed, 1 insertion(+) diff --git a/Mathlib/Data/List/Infix.lean b/Mathlib/Data/List/Infix.lean index b7bf8b65b8356..030675eb9f683 100644 --- a/Mathlib/Data/List/Infix.lean +++ b/Mathlib/Data/List/Infix.lean @@ -70,6 +70,7 @@ theorem mem_of_mem_dropLast (h : a ∈ l.dropLast) : a ∈ l := dropLast_subset l h attribute [gcongr] Sublist.drop +attribute [refl] prefix_refl suffix_refl infix_refl theorem concat_get_prefix {x y : List α} (h : x <+: y) (hl : x.length < y.length) : x ++ [y.get ⟨x.length, hl⟩] <+: y := by From bacc8b0efb93fde1befaddb5815315d489625176 Mon Sep 17 00:00:00 2001 From: sven-manthe Date: Fri, 25 Oct 2024 20:58:34 +0000 Subject: [PATCH 430/505] chore: fix name "Set.cou_inter_self_right_eq_coe" (#18222) The old name contained a typo. The new one is symmetric to the name of the following lemma in the same file. --- Mathlib/Data/Set/Subset.lean | 4 +++- 1 file changed, 3 insertions(+), 1 deletion(-) diff --git a/Mathlib/Data/Set/Subset.lean b/Mathlib/Data/Set/Subset.lean index 675ae793d1bc8..6672f997e1bd5 100644 --- a/Mathlib/Data/Set/Subset.lean +++ b/Mathlib/Data/Set/Subset.lean @@ -112,8 +112,10 @@ lemma image_val_union_self_left_eq : ↑D ∪ A = A := union_eq_right.2 image_val_subset @[simp] -lemma cou_inter_self_right_eq_coe : A ∩ ↑D = ↑D := +lemma image_val_inter_self_right_eq_coe : A ∩ ↑D = ↑D := inter_eq_right.2 image_val_subset +@[deprecated (since := "2024-10-25")] +alias cou_inter_self_right_eq_coe := image_val_inter_self_right_eq_coe @[simp] lemma image_val_inter_self_left_eq_coe : ↑D ∩ A = ↑D := From 70b8abbc768538873abfce8ca865634b1657b46e Mon Sep 17 00:00:00 2001 From: Bhavik Mehta Date: Fri, 25 Oct 2024 20:58:36 +0000 Subject: [PATCH 431/505] feat(Measure/Dirac): extensionality of measures by singletons (#18232) Two simple lemmas useful for discrete probability. --- Mathlib/MeasureTheory/Measure/Dirac.lean | 8 ++++++++ 1 file changed, 8 insertions(+) diff --git a/Mathlib/MeasureTheory/Measure/Dirac.lean b/Mathlib/MeasureTheory/Measure/Dirac.lean index 29e9d534230ce..bf8d888f3eea2 100644 --- a/Mathlib/MeasureTheory/Measure/Dirac.lean +++ b/Mathlib/MeasureTheory/Measure/Dirac.lean @@ -80,6 +80,14 @@ theorem restrict_singleton (μ : Measure α) (a : α) : μ.restrict {a} = μ {a} · have : s ∩ {a} = ∅ := inter_singleton_eq_empty.2 ha simp [*] +/-- Two measures on a countable space are equal if they agree on singletons. -/ +theorem ext_of_singleton [Countable α] {μ ν : Measure α} (h : ∀ a, μ {a} = ν {a}) : μ = ν := + ext_of_sUnion_eq_univ (countable_range singleton) (by aesop) (by aesop) + +/-- Two measures on a countable space are equal if and only if they agree on singletons. -/ +theorem ext_iff_singleton [Countable α] {μ ν : Measure α} : μ = ν ↔ ∀ a, μ {a} = ν {a} := + ⟨fun h _ ↦ h ▸ rfl, ext_of_singleton⟩ + /-- If `f` is a map with countable codomain, then `μ.map f` is a sum of Dirac measures. -/ theorem map_eq_sum [Countable β] [MeasurableSingletonClass β] (μ : Measure α) (f : α → β) (hf : Measurable f) : μ.map f = sum fun b : β => μ (f ⁻¹' {b}) • dirac b := by From 989ad277b63099209851afc2aded809cc5120abc Mon Sep 17 00:00:00 2001 From: Jakob von Raumer Date: Fri, 25 Oct 2024 21:52:01 +0000 Subject: [PATCH 432/505] feat(CategoryTheory/KanExtension): Composing `lan` with `colim` (#17355) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Joël Riou <37772949+joelriou@users.noreply.github.com> --- Mathlib/CategoryTheory/Functor/Const.lean | 7 ++ .../Functor/KanExtension/Adjunction.lean | 30 +++++- .../Functor/KanExtension/Basic.lean | 100 ++++++++++++++++++ 3 files changed, 136 insertions(+), 1 deletion(-) diff --git a/Mathlib/CategoryTheory/Functor/Const.lean b/Mathlib/CategoryTheory/Functor/Const.lean index 6d6844ebf3690..ba2d6a7b37e98 100644 --- a/Mathlib/CategoryTheory/Functor/Const.lean +++ b/Mathlib/CategoryTheory/Functor/Const.lean @@ -99,6 +99,13 @@ def compConstIso (F : C ⥤ D) : (fun X => NatIso.ofComponents (fun _ => Iso.refl _) (by aesop_cat)) (by aesop_cat) +/-- The canonical isomorphism +`const D ⋙ (whiskeringLeft J _ _).obj F ≅ const J`.-/ +@[simps!] +def constCompWhiskeringLeftIso (F : J ⥤ D) : + const D ⋙ (whiskeringLeft J D C).obj F ≅ const J := + NatIso.ofComponents fun X => NatIso.ofComponents fun Y => Iso.refl _ + end end CategoryTheory.Functor diff --git a/Mathlib/CategoryTheory/Functor/KanExtension/Adjunction.lean b/Mathlib/CategoryTheory/Functor/KanExtension/Adjunction.lean index 274273b8794c3..8998a80ca9e01 100644 --- a/Mathlib/CategoryTheory/Functor/KanExtension/Adjunction.lean +++ b/Mathlib/CategoryTheory/Functor/KanExtension/Adjunction.lean @@ -22,7 +22,7 @@ right Kan extension along `L`. namespace CategoryTheory -open Category +open Category Limits namespace Functor @@ -122,6 +122,20 @@ lemma isIso_lanAdjunction_counit_app_iff (G : D ⥤ H) : IsIso ((L.lanAdjunction H).counit.app G) ↔ G.IsLeftKanExtension (𝟙 (L ⋙ G)) := (isLeftKanExtension_iff_isIso _ (L.lanUnit.app (L ⋙ G)) _ (by simp)).symm +/-- Composing the left Kan extension of `L : C ⥤ D` with `colim` on shapes `D` is isomorphic +to `colim` on shapes `C`. -/ +@[simps!] +noncomputable def lanCompColimIso (L : C ⥤ D) [∀ (G : C ⥤ H), L.HasLeftKanExtension G] + [HasColimitsOfShape C H] [HasColimitsOfShape D H] : L.lan ⋙ colim ≅ colim (C := H) := + Iso.symm <| NatIso.ofComponents + (fun G ↦ (colimitIsoOfIsLeftKanExtension _ (L.lanUnit.app G)).symm) + (fun f ↦ colimit.hom_ext (fun i ↦ by + dsimp + rw [ι_colimMap_assoc, ι_colimitIsoOfIsLeftKanExtension_inv, + ι_colimitIsoOfIsLeftKanExtension_inv_assoc, ι_colimMap, ← assoc, ← assoc] + congr 1 + exact congr_app (L.lanUnit.naturality f) i)) + end section @@ -251,6 +265,20 @@ lemma isIso_ranAdjunction_unit_app_iff (G : D ⥤ H) : IsIso ((L.ranAdjunction H).unit.app G) ↔ G.IsRightKanExtension (𝟙 (L ⋙ G)) := (isRightKanExtension_iff_isIso _ (L.ranCounit.app (L ⋙ G)) _ (by simp)).symm +/-- Composing the right Kan extension of `L : C ⥤ D` with `lim` on shapes `D` is isomorphic +to `lim` on shapes `C`. -/ +@[simps!] +noncomputable def ranCompLimIso (L : C ⥤ D) [∀ (G : C ⥤ H), L.HasRightKanExtension G] + [HasLimitsOfShape C H] [HasLimitsOfShape D H] : L.ran ⋙ lim ≅ lim (C := H) := + NatIso.ofComponents + (fun G ↦ limitIsoOfIsRightKanExtension _ (L.ranCounit.app G)) + (fun f ↦ limit.hom_ext (fun i ↦ by + dsimp + rw [assoc, assoc, limMap_π, limitIsoOfIsRightKanExtension_hom_π_assoc, + limitIsoOfIsRightKanExtension_hom_π, limMap_π_assoc] + congr 1 + exact congr_app (L.ranCounit.naturality f) i)) + end section diff --git a/Mathlib/CategoryTheory/Functor/KanExtension/Basic.lean b/Mathlib/CategoryTheory/Functor/KanExtension/Basic.lean index 7c08315b07111..4ea68229bb26b 100644 --- a/Mathlib/CategoryTheory/Functor/KanExtension/Basic.lean +++ b/Mathlib/CategoryTheory/Functor/KanExtension/Basic.lean @@ -544,6 +544,106 @@ lemma isRightKanExtension_iff_of_iso₂ {F₁' F₂' : D ⥤ H} (α₁ : L ⋙ F end +section Colimit + +variable (F' : D ⥤ H) {L : C ⥤ D} {F : C ⥤ H} (α : F ⟶ L ⋙ F') [F'.IsLeftKanExtension α] + +/-- Construct a cocone for a left Kan extension `F' : D ⥤ H` of `F : C ⥤ H` along a functor +`L : C ⥤ D` given a cocone for `F`. -/ +@[simps] +noncomputable def coconeOfIsLeftKanExtension (c : Cocone F) : Cocone F' where + pt := c.pt + ι := F'.descOfIsLeftKanExtension α _ c.ι + +/-- If `c` is a colimit cocone for a functor `F : C ⥤ H` and `α : F ⟶ L ⋙ F'` is the unit of any +left Kan extension `F' : D ⥤ H` of `F` along `L : C ⥤ D`, then `coconeOfIsLeftKanExtension α c` is +a colimit cocone, too. -/ +@[simps] +def isColimitCoconeOfIsLeftKanExtension {c : Cocone F} (hc : IsColimit c) : + IsColimit (F'.coconeOfIsLeftKanExtension α c) where + desc s := hc.desc (Cocone.mk _ (α ≫ whiskerLeft L s.ι)) + fac s := by + have : F'.descOfIsLeftKanExtension α ((const D).obj c.pt) c.ι ≫ + (Functor.const _).map (hc.desc (Cocone.mk _ (α ≫ whiskerLeft L s.ι))) = s.ι := + F'.hom_ext_of_isLeftKanExtension α _ _ (by aesop_cat) + exact congr_app this + uniq s m hm := hc.hom_ext (fun j ↦ by + have := hm (L.obj j) + nth_rw 1 [← F'.descOfIsLeftKanExtension_fac_app α ((const D).obj c.pt)] + dsimp at this ⊢ + rw [assoc, this, IsColimit.fac, NatTrans.comp_app, whiskerLeft_app]) + +variable [HasColimit F] [HasColimit F'] + +/-- If `F' : D ⥤ H` is a left Kan extension of `F : C ⥤ H` along `L : C ⥤ D`, the colimit over `F'` +is isomorphic to the colimit over `F`. -/ +noncomputable def colimitIsoOfIsLeftKanExtension : colimit F' ≅ colimit F := + IsColimit.coconePointUniqueUpToIso (colimit.isColimit F') + (F'.isColimitCoconeOfIsLeftKanExtension α (colimit.isColimit F)) + +@[reassoc (attr := simp)] +lemma ι_colimitIsoOfIsLeftKanExtension_hom (i : C) : + α.app i ≫ colimit.ι F' (L.obj i) ≫ (F'.colimitIsoOfIsLeftKanExtension α).hom = + colimit.ι F i := by + simp [colimitIsoOfIsLeftKanExtension] + +@[reassoc (attr := simp)] +lemma ι_colimitIsoOfIsLeftKanExtension_inv (i : C) : + colimit.ι F i ≫ (F'.colimitIsoOfIsLeftKanExtension α).inv = + α.app i ≫ colimit.ι F' (L.obj i) := by + rw [Iso.comp_inv_eq, assoc, ι_colimitIsoOfIsLeftKanExtension_hom] + +end Colimit + +section Limit + +variable (F' : D ⥤ H) {L : C ⥤ D} {F : C ⥤ H} (α : L ⋙ F' ⟶ F) [F'.IsRightKanExtension α] + +/-- Construct a cone for a right Kan extension `F' : D ⥤ H` of `F : C ⥤ H` along a functor +`L : C ⥤ D` given a cone for `F`. -/ +@[simps] +noncomputable def coneOfIsRightKanExtension (c : Cone F) : Cone F' where + pt := c.pt + π := F'.liftOfIsRightKanExtension α _ c.π + +/-- If `c` is a limit cone for a functor `F : C ⥤ H` and `α : L ⋙ F' ⟶ F` is the counit of any +right Kan extension `F' : D ⥤ H` of `F` along `L : C ⥤ D`, then `coneOfIsRightKanExtension α c` is +a limit cone, too. -/ +@[simps] +def isLimitConeOfIsRightKanExtension {c : Cone F} (hc : IsLimit c) : + IsLimit (F'.coneOfIsRightKanExtension α c) where + lift s := hc.lift (Cone.mk _ (whiskerLeft L s.π ≫ α)) + fac s := by + have : (Functor.const _).map (hc.lift (Cone.mk _ (whiskerLeft L s.π ≫ α))) ≫ + F'.liftOfIsRightKanExtension α ((const D).obj c.pt) c.π = s.π := + F'.hom_ext_of_isRightKanExtension α _ _ (by aesop_cat) + exact congr_app this + uniq s m hm := hc.hom_ext (fun j ↦ by + have := hm (L.obj j) + nth_rw 1 [← F'.liftOfIsRightKanExtension_fac_app α ((const D).obj c.pt)] + dsimp at this ⊢ + rw [← assoc, this, IsLimit.fac, NatTrans.comp_app, whiskerLeft_app]) + +variable [HasLimit F] [HasLimit F'] + +/-- If `F' : D ⥤ H` is a right Kan extension of `F : C ⥤ H` along `L : C ⥤ D`, the limit over `F'` +is isomorphic to the limit over `F`. -/ +noncomputable def limitIsoOfIsRightKanExtension : limit F' ≅ limit F := + IsLimit.conePointUniqueUpToIso (limit.isLimit F') + (F'.isLimitConeOfIsRightKanExtension α (limit.isLimit F)) + +@[reassoc (attr := simp)] +lemma limitIsoOfIsRightKanExtension_inv_π (i : C) : + (F'.limitIsoOfIsRightKanExtension α).inv ≫ limit.π F' (L.obj i) ≫ α.app i = limit.π F i := by + simp [limitIsoOfIsRightKanExtension] + +@[reassoc (attr := simp)] +lemma limitIsoOfIsRightKanExtension_hom_π (i : C) : + (F'.limitIsoOfIsRightKanExtension α).hom ≫ limit.π F i = limit.π F' (L.obj i) ≫ α.app i := by + rw [← Iso.eq_inv_comp, limitIsoOfIsRightKanExtension_inv_π] + +end Limit + end Functor end CategoryTheory From 4d9990b3925435087a7e0530a4a1e84d7746c3e4 Mon Sep 17 00:00:00 2001 From: Jakob von Raumer Date: Fri, 25 Oct 2024 22:10:03 +0000 Subject: [PATCH 433/505] feat(CategoryTheory/Limits/Final): Add pointfree version of `Final.colimitIso` (#17902) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Joël Riou --- Mathlib/CategoryTheory/Limits/Final.lean | 59 +++++++++--------------- 1 file changed, 23 insertions(+), 36 deletions(-) diff --git a/Mathlib/CategoryTheory/Limits/Final.lean b/Mathlib/CategoryTheory/Limits/Final.lean index 8dea5b13733db..6390aaa372617 100644 --- a/Mathlib/CategoryTheory/Limits/Final.lean +++ b/Mathlib/CategoryTheory/Limits/Final.lean @@ -319,9 +319,21 @@ variable (G) https://stacks.math.columbia.edu/tag/04E7 -/ +@[simps! (config := .lemmasOnly)] def colimitIso [HasColimit G] : colimit (F ⋙ G) ≅ colimit G := asIso (colimit.pre G F) +/-- A pointfree version of `colimitIso`, stating that whiskering by `F` followed by taking the +colimit is isomorpic to taking the colimit on the codomain of `F`. -/ +def colimIso [HasColimitsOfShape D E] [HasColimitsOfShape C E] : + (whiskeringLeft _ _ _).obj F ⋙ colim ≅ colim (J := D) (C := E) := + NatIso.ofComponents (fun G => colimitIso F G) fun f => by + simp only [comp_obj, whiskeringLeft_obj_obj, colim_obj, comp_map, whiskeringLeft_obj_map, + colim_map, colimitIso_hom] + ext + simp only [comp_obj, ι_colimMap_assoc, whiskerLeft_app, colimit.ι_pre, colimit.ι_pre_assoc, + ι_colimMap] + end /-- Given a colimit cocone over `F ⋙ G` we can construct a colimit cocone over `G`. -/ @@ -342,24 +354,6 @@ include F in theorem hasColimitsOfShape_of_final [HasColimitsOfShape C E] : HasColimitsOfShape D E where has_colimit := fun _ => hasColimit_of_comp F -section - --- Porting note: this instance does not seem to be found automatically ---attribute [local instance] hasColimit_of_comp - -/-- When `F` is final, and `F ⋙ G` has a colimit, then `G` has a colimit also and -`colimit (F ⋙ G) ≅ colimit G` - -https://stacks.math.columbia.edu/tag/04E7 --/ -def colimitIso' [HasColimit (F ⋙ G)] : - haveI : HasColimit G := hasColimit_of_comp F - colimit (F ⋙ G) ≅ colimit G := - haveI : HasColimit G := hasColimit_of_comp F - asIso (colimit.pre G F) - -end - end Final end ArbitraryUniverse @@ -602,9 +596,20 @@ variable (G) https://stacks.math.columbia.edu/tag/04E7 -/ +@[simps! (config := .lemmasOnly)] def limitIso [HasLimit G] : limit (F ⋙ G) ≅ limit G := (asIso (limit.pre G F)).symm +/-- A pointfree version of `limitIso`, stating that whiskering by `F` followed by taking the +limit is isomorpic to taking the limit on the codomain of `F`. -/ +def limIso [HasLimitsOfShape D E] [HasLimitsOfShape C E] : + (whiskeringLeft _ _ _).obj F ⋙ lim ≅ lim (J := D) (C := E) := + Iso.symm <| NatIso.ofComponents (fun G => (limitIso F G).symm) fun f => by + simp only [comp_obj, whiskeringLeft_obj_obj, lim_obj, comp_map, whiskeringLeft_obj_map, lim_map, + Iso.symm_hom, limitIso_inv] + ext + simp + end /-- Given a limit cone over `F ⋙ G` we can construct a limit cone over `G`. -/ @@ -625,24 +630,6 @@ include F in theorem hasLimitsOfShape_of_initial [HasLimitsOfShape C E] : HasLimitsOfShape D E where has_limit := fun _ => hasLimit_of_comp F -section - --- Porting note: this instance does not seem to be found automatically --- attribute [local instance] hasLimit_of_comp - -/-- When `F` is initial, and `F ⋙ G` has a limit, then `G` has a limit also and -`limit (F ⋙ G) ≅ limit G` - -https://stacks.math.columbia.edu/tag/04E7 --/ -def limitIso' [HasLimit (F ⋙ G)] : - haveI : HasLimit G := hasLimit_of_comp F - limit (F ⋙ G) ≅ limit G := - haveI : HasLimit G := hasLimit_of_comp F - (asIso (limit.pre G F)).symm - -end - end Initial section From e91cdbcf5a387801a6225b41b0056c63206a9f7f Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sat, 26 Oct 2024 04:02:06 +0000 Subject: [PATCH 434/505] chore(ContinuousMap): move `compRight*` (#17935) Also drop some unneeded assumptions and rename/deprecate some lemmas. --- .../Instances.lean | 5 +- .../NonUnital.lean | 2 +- .../ContinuousFunctionalCalculus/Unique.lean | 8 +-- .../ContinuousFunctionalCalculus/Unital.lean | 2 +- .../Topology/Algebra/ContinuousMonoidHom.lean | 4 +- Mathlib/Topology/CompactOpen.lean | 46 ++++++++++++---- Mathlib/Topology/ContinuousMap/Algebra.lean | 4 ++ Mathlib/Topology/ContinuousMap/Compact.lean | 52 ------------------- Mathlib/Topology/ContinuousMap/Sigma.lean | 2 +- .../Topology/ContinuousMap/Weierstrass.lean | 2 +- Mathlib/Topology/Homotopy/HomotopyGroup.lean | 8 +-- 11 files changed, 55 insertions(+), 80 deletions(-) diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean index 1ed5d8debe5ce..b4cc643968f53 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean @@ -83,10 +83,7 @@ lemma isClosedEmbedding_cfcₙAux : IsClosedEmbedding (cfcₙAux hp₁ a ha) := refine ((cfcHom_isClosedEmbedding (hp₁.mpr ha)).comp ?_).comp ContinuousMapZero.isClosedEmbedding_toContinuousMap let e : C(σₙ 𝕜 a, 𝕜) ≃ₜ C(σ 𝕜 (a : A⁺¹), 𝕜) := - { (Homeomorph.compStarAlgEquiv' 𝕜 𝕜 <| .setCongr <| - (quasispectrum_eq_spectrum_inr' 𝕜 𝕜 a).symm) with - continuous_toFun := ContinuousMap.continuous_comp_left _ - continuous_invFun := ContinuousMap.continuous_comp_left _ } + (Homeomorph.setCongr (quasispectrum_eq_spectrum_inr' 𝕜 𝕜 a)).arrowCongr (.refl _) exact e.isClosedEmbedding @[deprecated (since := "2024-10-20")] diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean index ed76c4a88dc63..458c86fac7e39 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean @@ -178,7 +178,7 @@ theorem cfcₙHom_comp [UniqueNonUnitalContinuousFunctionalCalculus R A] (f : C( suffices cfcₙHom (cfcₙHom_predicate ha f) = φ from DFunLike.congr_fun this.symm g refine cfcₙHom_eq_of_continuous_of_map_id (cfcₙHom_predicate ha f) φ ?_ ?_ · refine (cfcₙHom_isClosedEmbedding ha).continuous.comp <| continuous_induced_rng.mpr ?_ - exact f'.toContinuousMap.continuous_comp_left.comp continuous_induced_dom + exact f'.toContinuousMap.continuous_precomp.comp continuous_induced_dom · simp only [φ, ψ, NonUnitalStarAlgHom.comp_apply, NonUnitalStarAlgHom.coe_mk', NonUnitalAlgHom.coe_mk] congr diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unique.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unique.lean index 5186931925074..0643b90789114 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unique.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unique.lean @@ -56,7 +56,7 @@ namespace ContinuousMap noncomputable def toNNReal (f : C(X, ℝ)) : C(X, ℝ≥0) := .realToNNReal |>.comp f @[fun_prop] -lemma continuous_toNNReal : Continuous (toNNReal (X := X)) := continuous_comp _ +lemma continuous_toNNReal : Continuous (toNNReal (X := X)) := continuous_postcomp _ @[simp] lemma toNNReal_apply (f : C(X, ℝ)) (x : X) : f.toNNReal x = (f x).toNNReal := rfl @@ -187,7 +187,7 @@ instance NNReal.instUniqueContinuousFunctionalCalculus [UniqueContinuousFunction Continuous ξ' ∧ ξ' (.restrict s' <| .id ℝ) = ξ (.restrict s <| .id ℝ≥0)) := by intro ξ' refine ⟨ξ.continuous_realContinuousMapOfNNReal hξ |>.comp <| - ContinuousMap.continuous_comp_left _, ?_⟩ + ContinuousMap.continuous_precomp _, ?_⟩ exact ξ.realContinuousMapOfNNReal_apply_comp_toReal (.restrict s <| .id ℝ≥0) obtain ⟨hφ', hφ_id⟩ := this φ hφ obtain ⟨hψ', hψ_id⟩ := this ψ hψ @@ -249,7 +249,7 @@ lemma toNNReal_apply (f : C(X, ℝ)₀) (x : X) : f.toNNReal x = Real.toNNReal ( lemma continuous_toNNReal : Continuous (toNNReal (X := X)) := by rw [continuous_induced_rng] convert_to Continuous (ContinuousMap.toNNReal ∘ ((↑) : C(X, ℝ)₀ → C(X, ℝ))) using 1 - exact ContinuousMap.continuous_comp _ |>.comp continuous_induced_dom + exact ContinuousMap.continuous_postcomp _ |>.comp continuous_induced_dom lemma toContinuousMapHom_toNNReal (f : C(X, ℝ)₀) : (toContinuousMapHom (X := X) (R := ℝ) f).toNNReal = @@ -401,7 +401,7 @@ instance NNReal.instUniqueNonUnitalContinuousFunctionalCalculus intro ξ' refine ⟨ξ.continuous_realContinuousMapZeroOfNNReal hξ |>.comp <| ?_, ?_⟩ · rw [continuous_induced_rng] - exact ContinuousMap.continuous_comp_left _ |>.comp continuous_induced_dom + exact ContinuousMap.continuous_precomp _ |>.comp continuous_induced_dom · exact ξ.realContinuousMapZeroOfNNReal_apply_comp_toReal (.id h0) obtain ⟨hφ', hφ_id⟩ := this φ hφ obtain ⟨hψ', hψ_id⟩ := this ψ hψ diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean index b193f3c40b439..820b45e9bdb3c 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean @@ -271,7 +271,7 @@ theorem cfcHom_comp [UniqueContinuousFunctionalCalculus R A] (f : C(spectrum R a (cfcHom ha).comp <| ContinuousMap.compStarAlgHom' R R f' suffices cfcHom (cfcHom_predicate ha f) = φ from DFunLike.congr_fun this.symm g refine cfcHom_eq_of_continuous_of_map_id (cfcHom_predicate ha f) φ ?_ ?_ - · exact (cfcHom_isClosedEmbedding ha).continuous.comp f'.continuous_comp_left + · exact (cfcHom_isClosedEmbedding ha).continuous.comp f'.continuous_precomp · simp only [φ, StarAlgHom.comp_apply, ContinuousMap.compStarAlgHom'_apply] congr ext x diff --git a/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean b/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean index faf9ab016a5e2..ff9ec26c16279 100644 --- a/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean +++ b/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean @@ -302,13 +302,13 @@ theorem continuous_comp [LocallyCompactSpace B] : theorem continuous_comp_left (f : ContinuousMonoidHom A B) : Continuous fun g : ContinuousMonoidHom B C => g.comp f := (inducing_toContinuousMap A C).continuous_iff.2 <| - f.toContinuousMap.continuous_comp_left.comp (inducing_toContinuousMap B C).continuous + f.toContinuousMap.continuous_precomp.comp (inducing_toContinuousMap B C).continuous @[to_additive] theorem continuous_comp_right (f : ContinuousMonoidHom B C) : Continuous fun g : ContinuousMonoidHom A B => f.comp g := (inducing_toContinuousMap A C).continuous_iff.2 <| - f.toContinuousMap.continuous_comp.comp (inducing_toContinuousMap A B).continuous + f.toContinuousMap.continuous_postcomp.comp (inducing_toContinuousMap A B).continuous variable (E) diff --git a/Mathlib/Topology/CompactOpen.lean b/Mathlib/Topology/CompactOpen.lean index 4ff883b2e49b1..8f58dc36f0245 100644 --- a/Mathlib/Topology/CompactOpen.lean +++ b/Mathlib/Topology/CompactOpen.lean @@ -78,27 +78,45 @@ lemma continuous_compactOpen {f : X → C(Y, Z)} : section Functorial /-- `C(X, ·)` is a functor. -/ -theorem continuous_comp (g : C(Y, Z)) : Continuous (ContinuousMap.comp g : C(X, Y) → C(X, Z)) := +theorem continuous_postcomp (g : C(Y, Z)) : Continuous (ContinuousMap.comp g : C(X, Y) → C(X, Z)) := continuous_compactOpen.2 fun _K hK _U hU ↦ isOpen_setOf_mapsTo hK (hU.preimage g.2) +@[deprecated (since := "2024-10-19")] alias continuous_comp := continuous_postcomp + /-- If `g : C(Y, Z)` is a topology inducing map, then the composition `ContinuousMap.comp g : C(X, Y) → C(X, Z)` is a topology inducing map too. -/ -theorem inducing_comp (g : C(Y, Z)) (hg : Inducing g) : Inducing (g.comp : C(X, Y) → C(X, Z)) where +theorem inducing_postcomp (g : C(Y, Z)) (hg : Inducing g) : + Inducing (g.comp : C(X, Y) → C(X, Z)) where induced := by simp only [compactOpen_eq, induced_generateFrom_eq, image_image2, hg.setOf_isOpen, image2_image_right, MapsTo, mem_preimage, preimage_setOf_eq, comp_apply] +@[deprecated (since := "2024-10-19")] alias inducing_comp := inducing_postcomp + /-- If `g : C(Y, Z)` is a topological embedding, then the composition `ContinuousMap.comp g : C(X, Y) → C(X, Z)` is an embedding too. -/ -theorem embedding_comp (g : C(Y, Z)) (hg : Embedding g) : Embedding (g.comp : C(X, Y) → C(X, Z)) := - ⟨inducing_comp g hg.1, fun _ _ ↦ (cancel_left hg.2).1⟩ +theorem embedding_postcomp (g : C(Y, Z)) (hg : Embedding g) : + Embedding (g.comp : C(X, Y) → C(X, Z)) := + ⟨inducing_postcomp g hg.1, fun _ _ ↦ (cancel_left hg.2).1⟩ + +@[deprecated (since := "2024-10-19")] alias embedding_comp := embedding_postcomp /-- `C(·, Z)` is a functor. -/ -@[fun_prop] -theorem continuous_comp_left (f : C(X, Y)) : Continuous (fun g => g.comp f : C(Y, Z) → C(X, Z)) := +@[continuity, fun_prop] +theorem continuous_precomp (f : C(X, Y)) : Continuous (fun g => g.comp f : C(Y, Z) → C(X, Z)) := continuous_compactOpen.2 fun K hK U hU ↦ by simpa only [mapsTo_image_iff] using isOpen_setOf_mapsTo (hK.image f.2) hU +@[deprecated (since := "2024-10-19")] alias continuous_comp_left := continuous_precomp + +variable (Z) in +/-- Precomposition by a continuous map is itself a continuous map between spaces of continuous maps. +-/ +@[simps apply] +def compRightContinuousMap (f : C(X, Y)) : + C(C(Y, Z), C(X, Z)) where + toFun g := g.comp f + /-- Any pair of homeomorphisms `X ≃ₜ Z` and `Y ≃ₜ T` gives rise to a homeomorphism `C(X, Y) ≃ₜ C(Z, T)`. -/ protected def _root_.Homeomorph.arrowCongr (φ : X ≃ₜ Z) (ψ : Y ≃ₜ T) : @@ -107,8 +125,16 @@ protected def _root_.Homeomorph.arrowCongr (φ : X ≃ₜ Z) (ψ : Y ≃ₜ T) : invFun f := .comp ψ.symm <| f.comp φ left_inv f := ext fun _ ↦ ψ.left_inv (f _) |>.trans <| congrArg f <| φ.left_inv _ right_inv f := ext fun _ ↦ ψ.right_inv (f _) |>.trans <| congrArg f <| φ.right_inv _ - continuous_toFun := continuous_comp _ |>.comp <| continuous_comp_left _ - continuous_invFun := continuous_comp _ |>.comp <| continuous_comp_left _ + continuous_toFun := continuous_postcomp _ |>.comp <| continuous_precomp _ + continuous_invFun := continuous_postcomp _ |>.comp <| continuous_precomp _ + +variable (Z) in +/-- Precomposition by a homeomorphism is itself a homeomorphism between spaces of continuous maps. +-/ +@[deprecated Homeomorph.arrowCongr (since := "2024-10-19")] +def compRightHomeomorph (f : X ≃ₜ Y) : + C(Y, Z) ≃ₜ C(X, Z) := + .arrowCongr f.symm (.refl _) variable [LocallyCompactPair Y Z] @@ -234,7 +260,7 @@ section InfInduced /-- For any subset `s` of `X`, the restriction of continuous functions to `s` is continuous as a function from `C(X, Y)` to `C(s, Y)` with their respective compact-open topologies. -/ theorem continuous_restrict (s : Set X) : Continuous fun F : C(X, Y) => F.restrict s := - continuous_comp_left <| restrict s <| .id X + continuous_precomp <| restrict s <| .id X theorem compactOpen_le_induced (s : Set X) : (ContinuousMap.compactOpen : TopologicalSpace C(X, Y)) ≤ @@ -340,7 +366,7 @@ section Curry compact, then this is a homeomorphism, see `Homeomorph.curry`. -/ def curry (f : C(X × Y, Z)) : C(X, C(Y, Z)) where toFun a := ⟨Function.curry f a, f.continuous.comp <| by fun_prop⟩ - continuous_toFun := (continuous_comp f).comp continuous_coev + continuous_toFun := (continuous_postcomp f).comp continuous_coev @[simp] theorem curry_apply (f : C(X × Y, Z)) (a : X) (b : Y) : f.curry a b = f (a, b) := diff --git a/Mathlib/Topology/ContinuousMap/Algebra.lean b/Mathlib/Topology/ContinuousMap/Algebra.lean index c475ae94f7dc9..c36b582bb8332 100644 --- a/Mathlib/Topology/ContinuousMap/Algebra.lean +++ b/Mathlib/Topology/ContinuousMap/Algebra.lean @@ -697,6 +697,10 @@ def ContinuousMap.compRightAlgHom {α β : Type*} [TopologicalSpace α] [Topolog map_mul' _ _ := ext fun _ ↦ rfl commutes' _ := ext fun _ ↦ rfl +theorem ContinuousMap.compRightAlgHom_continuous {α β : Type*} [TopologicalSpace α] + [TopologicalSpace β] (f : C(α, β)) : Continuous (compRightAlgHom R A f) := + continuous_precomp f + variable {A} /-- Coercion to a function as an `AlgHom`. -/ diff --git a/Mathlib/Topology/ContinuousMap/Compact.lean b/Mathlib/Topology/ContinuousMap/Compact.lean index 24fb53d8c00e8..9400fa18c4860 100644 --- a/Mathlib/Topology/ContinuousMap/Compact.lean +++ b/Mathlib/Topology/ContinuousMap/Compact.lean @@ -417,57 +417,6 @@ end CompLeft namespace ContinuousMap -/-! -We now setup variations on `compRight* f`, where `f : C(X, Y)` -(that is, precomposition by a continuous map), -as a morphism `C(Y, T) → C(X, T)`, respecting various types of structure. - -In particular: -* `compRightContinuousMap`, the bundled continuous map (for this we need `X Y` compact). -* `compRightHomeomorph`, when we precompose by a homeomorphism. -* `compRightAlgHom`, when `T = R` is a topological ring. --/ - - -section CompRight - -/-- Precomposition by a continuous map is itself a continuous map between spaces of continuous maps. --/ -def compRightContinuousMap {X Y : Type*} (T : Type*) [TopologicalSpace X] [CompactSpace X] - [TopologicalSpace Y] [CompactSpace Y] [PseudoMetricSpace T] (f : C(X, Y)) : - C(C(Y, T), C(X, T)) where - toFun g := g.comp f - continuous_toFun := by - refine Metric.continuous_iff.mpr ?_ - intro g ε ε_pos - refine ⟨ε, ε_pos, fun g' h => ?_⟩ - rw [ContinuousMap.dist_lt_iff ε_pos] at h ⊢ - exact fun x => h (f x) - -@[simp] -theorem compRightContinuousMap_apply {X Y : Type*} (T : Type*) [TopologicalSpace X] - [CompactSpace X] [TopologicalSpace Y] [CompactSpace Y] [PseudoMetricSpace T] (f : C(X, Y)) - (g : C(Y, T)) : (compRightContinuousMap T f) g = g.comp f := - rfl - -/-- Precomposition by a homeomorphism is itself a homeomorphism between spaces of continuous maps. --/ -def compRightHomeomorph {X Y : Type*} (T : Type*) [TopologicalSpace X] [CompactSpace X] - [TopologicalSpace Y] [CompactSpace Y] [PseudoMetricSpace T] (f : X ≃ₜ Y) : - C(Y, T) ≃ₜ C(X, T) where - toFun := compRightContinuousMap T f - invFun := compRightContinuousMap T f.symm - left_inv g := ext fun _ => congr_arg g (f.apply_symm_apply _) - right_inv g := ext fun _ => congr_arg g (f.symm_apply_apply _) - -theorem compRightAlgHom_continuous {X Y : Type*} (R A : Type*) [TopologicalSpace X] - [CompactSpace X] [TopologicalSpace Y] [CompactSpace Y] [CommSemiring R] [Semiring A] - [PseudoMetricSpace A] [TopologicalSemiring A] [Algebra R A] (f : C(X, Y)) : - Continuous (compRightAlgHom R A f) := - map_continuous (compRightContinuousMap A f) - -end CompRight - section LocalNormalConvergence /-! ### Local normal convergence @@ -476,7 +425,6 @@ A sum of continuous functions (on a locally compact space) is "locally normally sum of its sup-norms on any compact subset is summable. This implies convergence in the topology of `C(X, E)` (i.e. locally uniform convergence). -/ - open TopologicalSpace variable {X : Type*} [TopologicalSpace X] [LocallyCompactSpace X] diff --git a/Mathlib/Topology/ContinuousMap/Sigma.lean b/Mathlib/Topology/ContinuousMap/Sigma.lean index 320fd5cc4ab1f..f205d452cbc48 100644 --- a/Mathlib/Topology/ContinuousMap/Sigma.lean +++ b/Mathlib/Topology/ContinuousMap/Sigma.lean @@ -44,7 +44,7 @@ namespace ContinuousMap theorem embedding_sigmaMk_comp [Nonempty X] : Embedding (fun g : Σ i, C(X, Y i) ↦ (sigmaMk g.1).comp g.2) where toInducing := inducing_sigma.2 - ⟨fun i ↦ (sigmaMk i).inducing_comp embedding_sigmaMk.toInducing, fun i ↦ + ⟨fun i ↦ (sigmaMk i).inducing_postcomp embedding_sigmaMk.toInducing, fun i ↦ let ⟨x⟩ := ‹Nonempty X› ⟨_, (isOpen_sigma_fst_preimage {i}).preimage (continuous_eval_const x), fun _ ↦ Iff.rfl⟩⟩ inj := by diff --git a/Mathlib/Topology/ContinuousMap/Weierstrass.lean b/Mathlib/Topology/ContinuousMap/Weierstrass.lean index 8e903e39101e1..5e1edd95c1dc8 100644 --- a/Mathlib/Topology/ContinuousMap/Weierstrass.lean +++ b/Mathlib/Topology/ContinuousMap/Weierstrass.lean @@ -58,7 +58,7 @@ theorem polynomialFunctions_closure_eq_top (a b : ℝ) : compRightAlgHom ℝ ℝ (iccHomeoI a b h).symm -- This operation is itself a homeomorphism -- (with respect to the norm topologies on continuous functions). - let W' : C(Set.Icc a b, ℝ) ≃ₜ C(I, ℝ) := compRightHomeomorph ℝ (iccHomeoI a b h).symm + let W' : C(Set.Icc a b, ℝ) ≃ₜ C(I, ℝ) := (iccHomeoI a b h).arrowCongr (.refl _) have w : (W : C(Set.Icc a b, ℝ) → C(I, ℝ)) = W' := rfl -- Thus we take the statement of the Weierstrass approximation theorem for `[0,1]`, have p := polynomialFunctions_closure_eq_top' diff --git a/Mathlib/Topology/Homotopy/HomotopyGroup.lean b/Mathlib/Topology/Homotopy/HomotopyGroup.lean index 21ad316e36339..65ce882fec5b0 100644 --- a/Mathlib/Topology/Homotopy/HomotopyGroup.lean +++ b/Mathlib/Topology/Homotopy/HomotopyGroup.lean @@ -192,7 +192,7 @@ theorem continuous_toLoop (i : N) : Continuous (@toLoop N X _ x _ i) := (continuous_eval.comp <| Continuous.prodMap (ContinuousMap.continuous_curry.comp <| - (ContinuousMap.continuous_comp_left _).comp continuous_subtype_val) + (ContinuousMap.continuous_precomp _).comp continuous_subtype_val) continuous_id) _ @@ -211,9 +211,9 @@ def fromLoop (i : N) (p : Ω (Ω^ { j // j ≠ i } X x) const) : Ω^ N X x := · exact GenLoop.boundary _ _ ⟨⟨j, Hne⟩, Hj⟩⟩ theorem continuous_fromLoop (i : N) : Continuous (@fromLoop N X _ x _ i) := - ((ContinuousMap.continuous_comp_left _).comp <| + ((ContinuousMap.continuous_precomp _).comp <| ContinuousMap.continuous_uncurry.comp <| - (ContinuousMap.continuous_comp _).comp continuous_induced_dom).subtype_mk + (ContinuousMap.continuous_postcomp _).comp continuous_induced_dom).subtype_mk _ theorem to_from (i : N) (p : Ω (Ω^ { j // j ≠ i } X x) const) : toLoop i (fromLoop i p) = p := by @@ -244,7 +244,7 @@ theorem fromLoop_apply (i : N) {p : Ω (Ω^ { j // j ≠ i } X x) const} {t : I^ /-- Composition with `Cube.insertAt` as a continuous map. -/ abbrev cCompInsert (i : N) : C(C(I^N, X), C(I × I^{ j // j ≠ i }, X)) := ⟨fun f => f.comp (Cube.insertAt i), - (toContinuousMap <| Cube.insertAt i).continuous_comp_left⟩ + (toContinuousMap <| Cube.insertAt i).continuous_precomp⟩ /-- A homotopy between `n+1`-dimensional loops `p` and `q` constant on the boundary seen as a homotopy between two paths in the space of `n`-dimensional paths. -/ From fd050c2d54938c9a9b5b827fee930496295b1b0d Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Sat, 26 Oct 2024 11:49:25 +0000 Subject: [PATCH 435/505] chore: improve hypotheses to `ContinuousMap.induction_on` (#18116) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Instead of requiring `∀ f, p f → p (star f)`, we only require `p (star id)` (and note that `p id` is already a requirement, so there is no added proof burden). --- .../ContinuousMap/StoneWeierstrass.lean | 35 ++++++++++--------- 1 file changed, 19 insertions(+), 16 deletions(-) diff --git a/Mathlib/Topology/ContinuousMap/StoneWeierstrass.lean b/Mathlib/Topology/ContinuousMap/StoneWeierstrass.lean index 7d274f86f8aab..54ba6ba6ff1d1 100644 --- a/Mathlib/Topology/ContinuousMap/StoneWeierstrass.lean +++ b/Mathlib/Topology/ContinuousMap/StoneWeierstrass.lean @@ -425,29 +425,31 @@ theorem polynomialFunctions.starClosure_topologicalClosure {𝕜 : Type*} [RCLik @[elab_as_elim] theorem ContinuousMap.induction_on {𝕜 : Type*} [RCLike 𝕜] {s : Set 𝕜} {p : C(s, 𝕜) → Prop} (const : ∀ r, p (.const s r)) (id : p (.restrict s <| .id 𝕜)) + (star_id : p (star (.restrict s <| .id 𝕜))) (add : ∀ f g, p f → p g → p (f + g)) (mul : ∀ f g, p f → p g → p (f * g)) - (star : ∀ f, p f → p (star f)) (closure : (∀ f ∈ (polynomialFunctions s).starClosure, p f) → ∀ f, p f) (f : C(s, 𝕜)) : p f := by refine closure (fun f hf => ?_) f rw [polynomialFunctions.starClosure_eq_adjoin_X] at hf - induction hf using StarAlgebra.adjoin_induction with + induction hf using Algebra.adjoin_induction with | mem f hf => - simp only [toContinuousMapOnAlgHom_apply, Set.mem_singleton_iff] at hf - rwa [hf, toContinuousMapOn_X_eq_restrict_id] + simp only [Set.mem_union, Set.mem_singleton_iff, Set.mem_star] at hf + rw [star_eq_iff_star_eq, eq_comm (b := f)] at hf + obtain (rfl | rfl) := hf + all_goals simpa only [toContinuousMapOnAlgHom_apply, toContinuousMapOn_X_eq_restrict_id] | algebraMap r => exact const r | add _ _ _ _ hf hg => exact add _ _ hf hg | mul _ _ _ _ hf hg => exact mul _ _ hf hg - | star _ _ hf => exact star _ hf open Topology in @[elab_as_elim] theorem ContinuousMap.induction_on_of_compact {𝕜 : Type*} [RCLike 𝕜] {s : Set 𝕜} [CompactSpace s] {p : C(s, 𝕜) → Prop} (const : ∀ r, p (.const s r)) (id : p (.restrict s <| .id 𝕜)) + (star_id : p (star (.restrict s <| .id 𝕜))) (add : ∀ f g, p f → p g → p (f + g)) (mul : ∀ f g, p f → p g → p (f * g)) - (star : ∀ f, p f → p (star f)) (frequently : ∀ f, (∃ᶠ g in 𝓝 f, p g) → p f) (f : C(s, 𝕜)) : + (frequently : ∀ f, (∃ᶠ g in 𝓝 f, p g) → p f) (f : C(s, 𝕜)) : p f := by - refine f.induction_on const id add mul star fun h f ↦ frequently f ?_ + refine f.induction_on const id star_id add mul fun h f ↦ frequently f ?_ have := polynomialFunctions.starClosure_topologicalClosure s ▸ mem_top (x := f) rw [← SetLike.mem_coe, topologicalClosure_coe, mem_closure_iff_frequently] at this exact this.mp <| .of_forall h @@ -601,31 +603,32 @@ lemma ContinuousMapZero.adjoin_id_dense {s : Set 𝕜} [Zero s] (h0 : ((0 : s) : /-- An induction principle for `C(s, 𝕜)₀`. -/ @[elab_as_elim] lemma ContinuousMapZero.induction_on {s : Set 𝕜} [Zero s] (h0 : ((0 : s) : 𝕜) = 0) - {p : C(s, 𝕜)₀ → Prop} (zero : p 0) (id : p (.id h0)) + {p : C(s, 𝕜)₀ → Prop} (zero : p 0) (id : p (.id h0)) (star_id : p (star (.id h0))) (add : ∀ f g, p f → p g → p (f + g)) (mul : ∀ f g, p f → p g → p (f * g)) - (smul : ∀ (r : 𝕜) f, p f → p (r • f)) (star : ∀ f, p f → p (star f)) + (smul : ∀ (r : 𝕜) f, p f → p (r • f)) (closure : (∀ f ∈ adjoin 𝕜 {(.id h0 : C(s, 𝕜)₀)}, p f) → ∀ f, p f) (f : C(s, 𝕜)₀) : p f := by refine closure (fun f hf => ?_) f - induction hf using NonUnitalStarAlgebra.adjoin_induction with + induction hf using NonUnitalAlgebra.adjoin_induction with | mem f hf => - simp only [Set.mem_singleton_iff] at hf - rwa [hf] + simp only [Set.mem_union, Set.mem_singleton_iff, Set.mem_star] at hf + rw [star_eq_iff_star_eq, eq_comm (b := f)] at hf + obtain (rfl | rfl) := hf + all_goals assumption | zero => exact zero | add _ _ _ _ hf hg => exact add _ _ hf hg | mul _ _ _ _ hf hg => exact mul _ _ hf hg | smul _ _ _ hf => exact smul _ _ hf - | star _ _ hf => exact star _ hf open Topology in @[elab_as_elim] theorem ContinuousMapZero.induction_on_of_compact {s : Set 𝕜} [Zero s] (h0 : ((0 : s) : 𝕜) = 0) [CompactSpace s] {p : C(s, 𝕜)₀ → Prop} (zero : p 0) (id : p (.id h0)) - (add : ∀ f g, p f → p g → p (f + g)) (mul : ∀ f g, p f → p g → p (f * g)) - (smul : ∀ (r : 𝕜) f, p f → p (r • f)) (star : ∀ f, p f → p (star f)) + (star_id : p (star (.id h0))) (add : ∀ f g, p f → p g → p (f + g)) + (mul : ∀ f g, p f → p g → p (f * g)) (smul : ∀ (r : 𝕜) f, p f → p (r • f)) (frequently : ∀ f, (∃ᶠ g in 𝓝 f, p g) → p f) (f : C(s, 𝕜)₀) : p f := by - refine f.induction_on h0 zero id add mul smul star fun h f ↦ frequently f ?_ + refine f.induction_on h0 zero id star_id add mul smul fun h f ↦ frequently f ?_ have := (ContinuousMapZero.adjoin_id_dense h0).closure_eq ▸ Set.mem_univ (x := f) exact mem_closure_iff_frequently.mp this |>.mp <| .of_forall h From e6ba8e0a61b2e29769b333a35aa978c7d3e46e39 Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Sat, 26 Oct 2024 11:58:36 +0000 Subject: [PATCH 436/505] feat(NumberTheory/LSeries/Positivity): new file (#18031) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We use the newly added positivity results for holomorphic functions to show that the L-function associated to a sequence `a : ℕ → ℂ` such that `a 1 > 0` and `a n` is nonnegative real for all `n` takes positive real values on the real axis. From [EulerProducts](https://github.com/MichaelStollBayreuth/EulerProducts); this is needed to prove the non-vanishing at `s = 1` of Dirichlet L-functions of nontrivial quadratic characters. --- Mathlib.lean | 1 + .../Analysis/SpecialFunctions/Pow/Real.lean | 7 ++ Mathlib/NumberTheory/LSeries/Basic.lean | 14 +++ Mathlib/NumberTheory/LSeries/Positivity.lean | 109 ++++++++++++++++++ 4 files changed, 131 insertions(+) create mode 100644 Mathlib/NumberTheory/LSeries/Positivity.lean diff --git a/Mathlib.lean b/Mathlib.lean index 61df8bf2227e8..4ddb054c46958 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3582,6 +3582,7 @@ import Mathlib.NumberTheory.LSeries.HurwitzZetaOdd import Mathlib.NumberTheory.LSeries.HurwitzZetaValues import Mathlib.NumberTheory.LSeries.Linearity import Mathlib.NumberTheory.LSeries.MellinEqDirichlet +import Mathlib.NumberTheory.LSeries.Positivity import Mathlib.NumberTheory.LSeries.RiemannZeta import Mathlib.NumberTheory.LSeries.ZMod import Mathlib.NumberTheory.LegendreSymbol.AddCharacter diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean index abdc88a5f991e..a2e922b310ba9 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean @@ -997,6 +997,13 @@ lemma abs_cpow_inv_two_im (x : ℂ) : |(x ^ (2⁻¹ : ℂ)).im| = sqrt ((abs x - ← Real.sqrt_eq_rpow, _root_.abs_mul, _root_.abs_of_nonneg (sqrt_nonneg _), abs_sin_half, ← sqrt_mul (abs.nonneg _), ← mul_div_assoc, mul_sub, mul_one, abs_mul_cos_arg] +open scoped ComplexOrder in +lemma inv_natCast_cpow_ofReal_pos {n : ℕ} (hn : n ≠ 0) (x : ℝ) : + 0 < ((n : ℂ) ^ (x : ℂ))⁻¹ := by + refine RCLike.inv_pos_of_pos ?_ + rw [show (n : ℂ) ^ (x : ℂ) = (n : ℝ) ^ (x : ℂ) from rfl, ← ofReal_cpow n.cast_nonneg'] + positivity + end Complex section Tactics diff --git a/Mathlib/NumberTheory/LSeries/Basic.lean b/Mathlib/NumberTheory/LSeries/Basic.lean index 7221b06b0360f..99720f95d8532 100644 --- a/Mathlib/NumberTheory/LSeries/Basic.lean +++ b/Mathlib/NumberTheory/LSeries/Basic.lean @@ -111,6 +111,20 @@ lemma norm_term_le_of_re_le_re (f : ℕ → ℂ) {s s' : ℂ} (h : s.re ≤ s'.r next => rfl next hn => gcongr; exact Nat.one_le_cast.mpr <| Nat.one_le_iff_ne_zero.mpr hn +section positivity + +open scoped ComplexOrder + +lemma term_nonneg {a : ℕ → ℂ} {n : ℕ} (h : 0 ≤ a n) (x : ℝ) : 0 ≤ term a x n := by + rw [term_def] + split_ifs with hn + exacts [le_rfl, mul_nonneg h (inv_natCast_cpow_ofReal_pos hn x).le] + +lemma term_pos {a : ℕ → ℂ} {n : ℕ} (hn : n ≠ 0) (h : 0 < a n) (x : ℝ) : 0 < term a x n := by + simpa only [term_of_ne_zero hn] using mul_pos h <| inv_natCast_cpow_ofReal_pos hn x + +end positivity + end LSeries /-! diff --git a/Mathlib/NumberTheory/LSeries/Positivity.lean b/Mathlib/NumberTheory/LSeries/Positivity.lean new file mode 100644 index 0000000000000..b21893b9ed342 --- /dev/null +++ b/Mathlib/NumberTheory/LSeries/Positivity.lean @@ -0,0 +1,109 @@ +/- +Copyright (c) 2024 Michael Stoll. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Chris Birkbeck, David Loeffler, Michael Stoll +-/ +import Mathlib.Analysis.Complex.TaylorSeries +import Mathlib.Analysis.Complex.Positivity +import Mathlib.NumberTheory.ArithmeticFunction +import Mathlib.NumberTheory.LSeries.Deriv + +/-! +# Positivity of values of L-series + +The main results of this file are as follows. + +* If `a : ℕ → ℂ` takes nonnegative real values and `a 1 > 0`, then `L a x > 0` + when `x : ℝ` is in the open half-plane of absolute convergence; see + `LSeries.positive` and `ArithmeticFunction.LSeries_positive`. + +* If in addition the L_series of `a` agrees on some open right half-plane where it + converges with an entire function `f`, then `f` is positive on the real axis; + see `LSeries.positive_of_eq_differentiable` and + `ArithmeticFunction.LSeries_positive_of_eq_differentiable`. +-/ + +open scoped ComplexOrder + +open Complex + +namespace LSeries + +/-- If all values of a `ℂ`-valued arithmetic function are nonnegative reals and `x` is a +real number in the domain of absolute convergence, then the `n`th iterated derivative +of the associated L-series is nonnegative real when `n` is even and nonpositive real +when `n` is odd. -/ +lemma iteratedDeriv_alternating {a : ℕ → ℂ} (hn : 0 ≤ a) {x : ℝ} + (h : LSeries.abscissaOfAbsConv a < x) (n : ℕ) : + 0 ≤ (-1) ^ n * iteratedDeriv n (LSeries a) x := by + rw [LSeries_iteratedDeriv _ h, LSeries, ← mul_assoc, ← pow_add, Even.neg_one_pow ⟨n, rfl⟩, + one_mul] + refine tsum_nonneg fun k ↦ ?_ + rw [LSeries.term_def] + split + · exact le_rfl + · refine mul_nonneg ?_ <| (inv_natCast_cpow_ofReal_pos (by assumption) x).le + induction n with + | zero => simpa only [Function.iterate_zero, id_eq] using hn k + | succ n IH => + rw [Function.iterate_succ_apply'] + refine mul_nonneg ?_ IH + simp only [← natCast_log, zero_le_real, Real.log_natCast_nonneg] + +/-- If all values of `a : ℕ → ℂ` are nonnegative reals and `a 1` is positive, +then `L a x` is positive real for all real `x` larger than `abscissaOfAbsConv a`. -/ +lemma positive {a : ℕ → ℂ} (ha₀ : 0 ≤ a) (ha₁ : 0 < a 1) {x : ℝ} (hx : abscissaOfAbsConv a < x) : + 0 < LSeries a x := by + rw [LSeries] + refine tsum_pos ?_ (fun n ↦ term_nonneg (ha₀ n) x) 1 <| term_pos one_ne_zero ha₁ x + exact LSeriesSummable_of_abscissaOfAbsConv_lt_re <| by simpa only [ofReal_re] using hx + +/-- If all values of `a : ℕ → ℂ` are nonnegative reals and `a 1` +is positive, and the L-series of `a` agrees with an entire function `f` on some open +right half-plane where it converges, then `f` is real and positive on `ℝ`. -/ +lemma positive_of_differentiable_of_eqOn {a : ℕ → ℂ} (ha₀ : 0 ≤ a) (ha₁ : 0 < a 1) {f : ℂ → ℂ} + (hf : Differentiable ℂ f) {x : ℝ} (hx : abscissaOfAbsConv a ≤ x) + (hf' : {s | x < s.re}.EqOn f (LSeries a)) (y : ℝ) : + 0 < f y := by + have hxy : x < max x y + 1 := (le_max_left x y).trans_lt (lt_add_one _) + have hxy' : abscissaOfAbsConv a < max x y + 1 := hx.trans_lt <| mod_cast hxy + have hys : (max x y + 1 : ℂ) ∈ {s | x < s.re} := by + simp only [Set.mem_setOf_eq, add_re, ofReal_re, one_re, hxy] + have hfx : 0 < f (max x y + 1) := by + simpa only [hf' hys, ofReal_add, ofReal_one] using positive ha₀ ha₁ hxy' + refine (hfx.trans_le <| hf.apply_le_of_iteratedDeriv_alternating (fun n _ ↦ ?_) ?_) + · have hs : IsOpen {s : ℂ | x < s.re} := continuous_re.isOpen_preimage _ isOpen_Ioi + simpa only [hf'.iteratedDeriv_of_isOpen hs n hys, ofReal_add, ofReal_one] using + iteratedDeriv_alternating ha₀ hxy' n + · exact_mod_cast (le_max_right x y).trans (lt_add_one _).le + +end LSeries + +namespace ArithmeticFunction + +/-- If all values of a `ℂ`-valued arithmetic function are nonnegative reals and `x` is a +real number in the domain of absolute convergence, then the `n`th iterated derivative +of the associated L-series is nonnegative real when `n` is even and nonpositive real +when `n` is odd. -/ +lemma iteratedDeriv_LSeries_alternating (a : ArithmeticFunction ℂ) (hn : ∀ n, 0 ≤ a n) {x : ℝ} + (h : LSeries.abscissaOfAbsConv a < x) (n : ℕ) : + 0 ≤ (-1) ^ n * iteratedDeriv n (LSeries (a ·)) x := + LSeries.iteratedDeriv_alternating hn h n + +/-- If all values of a `ℂ`-valued arithmetic function `a` are nonnegative reals and `a 1` is +positive, then `L a x` is positive real for all real `x` larger than `abscissaOfAbsConv a`. -/ +lemma LSeries_positive {a : ℕ → ℂ} (ha₀ : 0 ≤ a) (ha₁ : 0 < a 1) {x : ℝ} + (hx : LSeries.abscissaOfAbsConv a < x) : + 0 < LSeries a x := + LSeries.positive ha₀ ha₁ hx + +/-- If all values of a `ℂ`-valued arithmetic function `a` are nonnegative reals and `a 1` +is positive, and the L-series of `a` agrees with an entire function `f` on some open +right half-plane where it converges, then `f` is real and positive on `ℝ`. -/ +lemma LSeries_positive_of_differentiable_of_eqOn {a : ArithmeticFunction ℂ} (ha₀ : 0 ≤ (a ·)) + (ha₁ : 0 < a 1) {f : ℂ → ℂ} (hf : Differentiable ℂ f) {x : ℝ} + (hx : LSeries.abscissaOfAbsConv a ≤ x) (hf' : {s | x < s.re}.EqOn f (LSeries a)) (y : ℝ) : + 0 < f y := + LSeries.positive_of_differentiable_of_eqOn ha₀ ha₁ hf hx hf' y + +end ArithmeticFunction From 83461cc0d7b87880c93523b2ae1df94508776b53 Mon Sep 17 00:00:00 2001 From: Riccardo Brasca Date: Sat, 26 Oct 2024 11:58:37 +0000 Subject: [PATCH 437/505] feat: miscellaneous results about inertia degree (#18111) From flt-regular. --- Mathlib/NumberTheory/RamificationInertia.lean | 37 +++++++++++++++++++ Mathlib/RingTheory/Ideal/Maps.lean | 2 + 2 files changed, 39 insertions(+) diff --git a/Mathlib/NumberTheory/RamificationInertia.lean b/Mathlib/NumberTheory/RamificationInertia.lean index 71fc4378b8d5f..cd2da449ce790 100644 --- a/Mathlib/NumberTheory/RamificationInertia.lean +++ b/Mathlib/NumberTheory/RamificationInertia.lean @@ -124,6 +124,19 @@ theorem le_comap_pow_ramificationIdx : p ≤ comap f (P ^ ramificationIdx f p P) theorem le_comap_of_ramificationIdx_ne_zero (h : ramificationIdx f p P ≠ 0) : p ≤ comap f P := Ideal.map_le_iff_le_comap.mp <| le_pow_ramificationIdx.trans <| Ideal.pow_le_self <| h +variable {S₁: Type*} [CommRing S₁] [Algebra R S₁] + +lemma ramificationIdx_comap_eq [Algebra R S] (e : S ≃ₐ[R] S₁) (P : Ideal S₁) : + ramificationIdx (algebraMap R S) p (P.comap e) = ramificationIdx (algebraMap R S₁) p P := by + dsimp only [ramificationIdx] + congr + ext n + simp only [Set.mem_setOf_eq, Ideal.map_le_iff_le_comap] + rw [← comap_coe e, ← e.toRingEquiv_toRingHom, comap_coe, ← RingEquiv.symm_symm (e : S ≃+* S₁), + ← map_comap_of_equiv, ← Ideal.map_pow, map_comap_of_equiv, ← comap_coe (RingEquiv.symm _), + comap_comap, RingEquiv.symm_symm, e.toRingEquiv_toRingHom, ← e.toAlgHom_toRingHom, + AlgHom.comp_algebraMap] + namespace IsDedekindDomain variable [IsDedekindDomain S] @@ -202,6 +215,30 @@ theorem inertiaDeg_algebraMap [Algebra R S] [Algebra (R ⧸ p) (S ⧸ P)] rw [Ideal.Quotient.lift_mk, ← Ideal.Quotient.algebraMap_eq P, ← IsScalarTower.algebraMap_eq, ← Ideal.Quotient.algebraMap_eq, ← IsScalarTower.algebraMap_apply] +lemma inertiaDeg_comap_eq [Algebra R S] (e : S ≃ₐ[R] S₁) (P : Ideal S₁) [p.IsMaximal] : + inertiaDeg (algebraMap R S) p (P.comap e) = inertiaDeg (algebraMap R S₁) p P := by + dsimp only [Ideal.inertiaDeg] + have : (P.comap e).comap (algebraMap R S) = p ↔ P.comap (algebraMap R S₁) = p := by + rw [← comap_coe e, comap_comap, ← e.toAlgHom_toRingHom, AlgHom.comp_algebraMap] + split + swap + · rw [dif_neg]; rwa [← this] + rw [dif_pos (this.mp ‹_›)] + apply (config := { allowSynthFailures := true }) LinearEquiv.finrank_eq + have E := quotientEquivAlg (P.comap e) P e (map_comap_of_surjective _ e.surjective P).symm + apply (config := {allowSynthFailures := true }) LinearEquiv.mk + case toLinearMap => + apply (config := {allowSynthFailures := true }) LinearMap.mk + swap + · exact E.toLinearMap.toAddHom + · intro r x + show E (_ * _) = _ * (E x) + obtain ⟨r, rfl⟩ := Ideal.Quotient.mk_surjective r + simp only [Quotient.mk_comp_algebraMap, Quotient.lift_mk, _root_.map_mul, AlgEquiv.commutes, + RingHomCompTriple.comp_apply] + · exact E.left_inv + · exact E.right_inv + end DecEq section FinrankQuotientMap diff --git a/Mathlib/RingTheory/Ideal/Maps.lean b/Mathlib/RingTheory/Ideal/Maps.lean index 0462faa6fbffb..725382c711dc9 100644 --- a/Mathlib/RingTheory/Ideal/Maps.lean +++ b/Mathlib/RingTheory/Ideal/Maps.lean @@ -51,6 +51,8 @@ def comap [RingHomClass F R S] (I : Ideal S) : Ideal R where @[simp] theorem coe_comap [RingHomClass F R S] (I : Ideal S) : (comap f I : Set R) = f ⁻¹' I := rfl +lemma comap_coe [RingHomClass F R S] (I : Ideal S) : I.comap (f : R →+* S) = I.comap f := rfl + variable {f} theorem map_mono (h : I ≤ J) : map f I ≤ map f J := From 97a5c4677a62fce8992b5c438e12815686b3e50a Mon Sep 17 00:00:00 2001 From: Xavier Roblot Date: Sat, 26 Oct 2024 11:58:38 +0000 Subject: [PATCH 438/505] feat(Complex/Order): add some results about multiplication of a real and a complex (#18136) This PR is part of the proof of the Analytic Class Number Formula. Co-authored-by: Xavier Roblot <46200072+xroblot@users.noreply.github.com> --- Mathlib/Analysis/RCLike/Basic.lean | 14 ++++++++++++++ Mathlib/Data/Complex/Order.lean | 6 ++++++ 2 files changed, 20 insertions(+) diff --git a/Mathlib/Analysis/RCLike/Basic.lean b/Mathlib/Analysis/RCLike/Basic.lean index 6df3640fd74f1..039e769c1b674 100644 --- a/Mathlib/Analysis/RCLike/Basic.lean +++ b/Mathlib/Analysis/RCLike/Basic.lean @@ -872,6 +872,20 @@ instance {A : Type*} [NonUnitalRing A] [StarRing A] [PartialOrder A] [StarOrdere scoped[ComplexOrder] attribute [instance] StarModule.instOrderedSMul +theorem ofReal_mul_pos_iff (x : ℝ) (z : K) : + 0 < x * z ↔ (x < 0 ∧ z < 0) ∨ (0 < x ∧ 0 < z) := by + simp only [pos_iff (K := K), neg_iff (K := K), re_ofReal_mul, im_ofReal_mul] + obtain hx | hx | hx := lt_trichotomy x 0 + · simp only [mul_pos_iff, not_lt_of_gt hx, false_and, hx, true_and, false_or, mul_eq_zero, hx.ne, + or_false] + · simp only [hx, zero_mul, lt_self_iff_false, false_and, false_or] + · simp only [mul_pos_iff, hx, true_and, not_lt_of_gt hx, false_and, or_false, mul_eq_zero, + hx.ne', false_or] + +theorem ofReal_mul_neg_iff (x : ℝ) (z : K) : + x * z < 0 ↔ (x < 0 ∧ 0 < z) ∨ (0 < x ∧ z < 0) := by + simpa only [mul_neg, neg_pos, neg_neg_iff_pos] using ofReal_mul_pos_iff x (-z) + end Order section CleanupLemmas diff --git a/Mathlib/Data/Complex/Order.lean b/Mathlib/Data/Complex/Order.lean index 47d0dd85378f3..9e641fe376939 100644 --- a/Mathlib/Data/Complex/Order.lean +++ b/Mathlib/Data/Complex/Order.lean @@ -63,6 +63,12 @@ theorem nonneg_iff {z : ℂ} : 0 ≤ z ↔ 0 ≤ z.re ∧ 0 = z.im := theorem pos_iff {z : ℂ} : 0 < z ↔ 0 < z.re ∧ 0 = z.im := lt_def +theorem nonpos_iff {z : ℂ} : z ≤ 0 ↔ z.re ≤ 0 ∧ z.im = 0 := + le_def + +theorem neg_iff {z : ℂ} : z < 0 ↔ z.re < 0 ∧ z.im = 0 := + lt_def + @[simp, norm_cast] theorem real_le_real {x y : ℝ} : (x : ℂ) ≤ (y : ℂ) ↔ x ≤ y := by simp [le_def, ofReal] From 85ae613a6bf0c17553a1f50756c49f8d17decb2a Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Sat, 26 Oct 2024 11:58:39 +0000 Subject: [PATCH 439/505] feat(Algebra/BigOperators/GroupWithZero/Action): Action of group permutes a product (#18148) This PR adds a couple lemmas for when the action of a group permutes a product. Co-authored-by: Thomas Browning --- .../Algebra/BigOperators/GroupWithZero/Action.lean | 12 ++++++++++++ 1 file changed, 12 insertions(+) diff --git a/Mathlib/Algebra/BigOperators/GroupWithZero/Action.lean b/Mathlib/Algebra/BigOperators/GroupWithZero/Action.lean index c969334d05c0a..6f377d4ff77be 100644 --- a/Mathlib/Algebra/BigOperators/GroupWithZero/Action.lean +++ b/Mathlib/Algebra/BigOperators/GroupWithZero/Action.lean @@ -73,6 +73,18 @@ theorem smul_finprod' {ι : Sort*} [Finite ι] {f : ι → β} (r : α) : simp only [finprod_eq_prod_plift_of_mulSupport_subset (s := Finset.univ) (by simp), finprod_eq_prod_of_fintype, Finset.smul_prod'] +variable {G : Type*} [Group G] [MulDistribMulAction G β] + +theorem Finset.smul_prod_perm [Fintype G] (b : β) (g : G) : + (g • ∏ h : G, h • b) = ∏ h : G, h • b := by + simp only [smul_prod', smul_smul] + exact Finset.prod_bijective (g * ·) (Group.mulLeft_bijective g) (by simp) (fun _ _ ↦ rfl) + +theorem smul_finprod_perm [Finite G] (b : β) (g : G) : + (g • ∏ᶠ h : G, h • b) = ∏ᶠ h : G, h • b := by + cases nonempty_fintype G + simp only [finprod_eq_prod_of_fintype, Finset.smul_prod_perm] + end namespace List From a0de06886e296efb09c13ac6af1066a06c2be252 Mon Sep 17 00:00:00 2001 From: Xavier Roblot Date: Sat, 26 Oct 2024 11:58:41 +0000 Subject: [PATCH 440/505] feat(LSeries/RiemannZeta): versions of the residue computation of the zeta function with tsum and for a real variable (#18166) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The proof of: ``` Tendsto (fun s : ℝ ↦ (s - 1) * ∑' (n : ℕ), 1 / (n : ℝ) ^ s) (𝓝[>] 1) (𝓝 1) ``` using the existing result for the complex variable. This PR is part of the proof of the Analytic Class Number Formula. --- Mathlib/NumberTheory/LSeries/RiemannZeta.lean | 19 +++++++++++++++++++ 1 file changed, 19 insertions(+) diff --git a/Mathlib/NumberTheory/LSeries/RiemannZeta.lean b/Mathlib/NumberTheory/LSeries/RiemannZeta.lean index bcc35496bf6f6..0b29f4b14b30a 100644 --- a/Mathlib/NumberTheory/LSeries/RiemannZeta.lean +++ b/Mathlib/NumberTheory/LSeries/RiemannZeta.lean @@ -202,6 +202,25 @@ theorem zeta_nat_eq_tsum_of_gt_one {k : ℕ} (hk : 1 < k) : lemma riemannZeta_residue_one : Tendsto (fun s ↦ (s - 1) * riemannZeta s) (𝓝[≠] 1) (𝓝 1) := by exact hurwitzZetaEven_residue_one 0 +/-- The residue of `ζ(s)` at `s = 1` is equal to 1, expressed using `tsum`. -/ +theorem tendsto_sub_mul_tsum_nat_cpow : + Tendsto (fun s : ℂ ↦ (s - 1) * ∑' (n : ℕ), 1 / (n : ℂ) ^ s) (𝓝[{s | 1 < re s}] 1) (𝓝 1) := by + refine (tendsto_nhdsWithin_mono_left ?_ riemannZeta_residue_one).congr' ?_ + · simp only [subset_compl_singleton_iff, mem_setOf_eq, one_re, not_lt, le_refl] + · filter_upwards [eventually_mem_nhdsWithin] with s hs using + congr_arg _ <| zeta_eq_tsum_one_div_nat_cpow hs + +/-- The residue of `ζ(s)` at `s = 1` is equal to 1 expressed using `tsum` and for a +real variable. -/ +theorem tendsto_sub_mul_tsum_nat_rpow : + Tendsto (fun s : ℝ ↦ (s - 1) * ∑' (n : ℕ), 1 / (n : ℝ) ^ s) (𝓝[>] 1) (𝓝 1) := by + rw [← tendsto_ofReal_iff, ofReal_one] + have : Tendsto (fun s : ℝ ↦ (s : ℂ)) (𝓝[>] 1) (𝓝[{s | 1 < re s}] 1) := + continuous_ofReal.continuousWithinAt.tendsto_nhdsWithin (fun _ _ ↦ by aesop) + apply (tendsto_sub_mul_tsum_nat_cpow.comp this).congr fun s ↦ ?_ + simp only [one_div, Function.comp_apply, ofReal_mul, ofReal_sub, ofReal_one, ofReal_tsum, + ofReal_inv, ofReal_cpow (Nat.cast_nonneg _), ofReal_natCast] + /- naming scheme was changed from `riemannCompletedZeta` to `completedRiemannZeta`; add aliases for the old names -/ section aliases From 2e87d1a8a668a70fa5b9f24085270036824c1902 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sat, 26 Oct 2024 11:58:42 +0000 Subject: [PATCH 441/505] =?UTF-8?q?feat:=20`StarOrderedRing=20(=CE=B1=20?= =?UTF-8?q?=C3=97=20=CE=B2)`=20(#18209)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The `Pi` version (for `Finite` indices) would be more work due to missing API around `Submonoid.pi`. --- Mathlib.lean | 1 + Mathlib/Algebra/Group/Subgroup/Basic.lean | 9 ++++++ .../Algebra/Group/Submonoid/Operations.lean | 9 ++++++ Mathlib/Algebra/Order/Star/Prod.lean | 30 +++++++++++++++++++ Mathlib/Logic/Basic.lean | 4 +++ 5 files changed, 53 insertions(+) create mode 100644 Mathlib/Algebra/Order/Star/Prod.lean diff --git a/Mathlib.lean b/Mathlib.lean index 4ddb054c46958..81e1f700393e7 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -710,6 +710,7 @@ import Mathlib.Algebra.Order.Ring.Unbundled.Rat import Mathlib.Algebra.Order.Ring.WithTop import Mathlib.Algebra.Order.Star.Basic import Mathlib.Algebra.Order.Star.Conjneg +import Mathlib.Algebra.Order.Star.Prod import Mathlib.Algebra.Order.Sub.Basic import Mathlib.Algebra.Order.Sub.Defs import Mathlib.Algebra.Order.Sub.Prod diff --git a/Mathlib/Algebra/Group/Subgroup/Basic.lean b/Mathlib/Algebra/Group/Subgroup/Basic.lean index 9f971900d0349..ca7a9f590871e 100644 --- a/Mathlib/Algebra/Group/Subgroup/Basic.lean +++ b/Mathlib/Algebra/Group/Subgroup/Basic.lean @@ -1401,6 +1401,15 @@ theorem prod_le_iff {H : Subgroup G} {K : Subgroup N} {J : Subgroup (G × N)} : theorem prod_eq_bot_iff {H : Subgroup G} {K : Subgroup N} : H.prod K = ⊥ ↔ H = ⊥ ∧ K = ⊥ := by simpa only [← Subgroup.toSubmonoid_eq] using Submonoid.prod_eq_bot_iff +@[to_additive closure_prod] +theorem closure_prod {s : Set G} {t : Set N} (hs : 1 ∈ s) (ht : 1 ∈ t) : + closure (s ×ˢ t) = (closure s).prod (closure t) := + le_antisymm + (closure_le _ |>.2 <| Set.prod_subset_prod_iff.2 <| .inl ⟨subset_closure, subset_closure⟩) + (prod_le_iff.2 ⟨ + map_le_iff_le_comap.2 <| closure_le _ |>.2 fun _x hx => subset_closure ⟨hx, ht⟩, + map_le_iff_le_comap.2 <| closure_le _ |>.2 fun _y hy => subset_closure ⟨hs, hy⟩⟩) + /-- Product of subgroups is isomorphic to their product as groups. -/ @[to_additive prodEquiv "Product of additive subgroups is isomorphic to their product diff --git a/Mathlib/Algebra/Group/Submonoid/Operations.lean b/Mathlib/Algebra/Group/Submonoid/Operations.lean index e5ca00f78346e..ee672456df581 100644 --- a/Mathlib/Algebra/Group/Submonoid/Operations.lean +++ b/Mathlib/Algebra/Group/Submonoid/Operations.lean @@ -715,6 +715,15 @@ theorem prod_le_iff {s : Submonoid M} {t : Submonoid N} {u : Submonoid (M × N)} simpa using h2 simpa using Submonoid.mul_mem _ h1' h2' +@[to_additive closure_prod] +theorem closure_prod {s : Set M} {t : Set N} (hs : 1 ∈ s) (ht : 1 ∈ t) : + closure (s ×ˢ t) = (closure s).prod (closure t) := + le_antisymm + (closure_le.2 <| Set.prod_subset_prod_iff.2 <| .inl ⟨subset_closure, subset_closure⟩) + (prod_le_iff.2 ⟨ + map_le_of_le_comap _ <| closure_le.2 fun _x hx => subset_closure ⟨hx, ht⟩, + map_le_of_le_comap _ <| closure_le.2 fun _y hy => subset_closure ⟨hs, hy⟩⟩) + end Submonoid namespace MonoidHom diff --git a/Mathlib/Algebra/Order/Star/Prod.lean b/Mathlib/Algebra/Order/Star/Prod.lean new file mode 100644 index 0000000000000..f2375c7d57a81 --- /dev/null +++ b/Mathlib/Algebra/Order/Star/Prod.lean @@ -0,0 +1,30 @@ +/- +Copyright (c) 2024 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser +-/ +import Mathlib.Algebra.Order.Star.Basic +import Mathlib.Algebra.Star.Prod +import Mathlib.Algebra.Ring.Prod + +/-! +# Products of star-ordered rings +-/ + +variable {α β : Type*} + +open AddSubmonoid in +instance Prod.instStarOrderedRing + [NonUnitalSemiring α] [NonUnitalSemiring β] [PartialOrder α] [PartialOrder β] + [StarRing α] [StarRing β] [StarOrderedRing α] [StarOrderedRing β] : + StarOrderedRing (α × β) where + le_iff := Prod.forall.2 fun xa xy => Prod.forall.2 fun ya yb => by + have : + closure (Set.range fun s : α × β ↦ star s * s) = + (closure <| Set.range fun s : α ↦ star s * s).prod + (closure <| Set.range fun s : β ↦ star s * s) := by + rw [← closure_prod (Set.mem_range.2 ⟨0, by simp⟩) (Set.mem_range.2 ⟨0, by simp⟩), + Set.prod_range_range_eq] + simp_rw [Prod.mul_def, Prod.star_def] + simp only [mk_le_mk, Prod.exists, mk_add_mk, mk.injEq, StarOrderedRing.le_iff, this, + AddSubmonoid.mem_prod, exists_and_exists_comm, and_and_and_comm] diff --git a/Mathlib/Logic/Basic.lean b/Mathlib/Logic/Basic.lean index b4258e60f795c..cf41e0848cfdf 100644 --- a/Mathlib/Logic/Basic.lean +++ b/Mathlib/Logic/Basic.lean @@ -485,6 +485,10 @@ theorem imp_forall_iff {α : Type*} {p : Prop} {q : α → Prop} : (p → ∀ x, theorem exists_swap {p : α → β → Prop} : (∃ x y, p x y) ↔ ∃ y x, p x y := ⟨fun ⟨x, y, h⟩ ↦ ⟨y, x, h⟩, fun ⟨y, x, h⟩ ↦ ⟨x, y, h⟩⟩ +theorem exists_and_exists_comm {P : α → Prop} {Q : β → Prop} : + (∃ a, P a) ∧ (∃ b, Q b) ↔ ∃ a b, P a ∧ Q b := + ⟨fun ⟨⟨a, ha⟩, ⟨b, hb⟩⟩ ↦ ⟨a, b, ⟨ha, hb⟩⟩, fun ⟨a, b, ⟨ha, hb⟩⟩ ↦ ⟨⟨a, ha⟩, ⟨b, hb⟩⟩⟩ + export Classical (not_forall) theorem not_forall_not : (¬∀ x, ¬p x) ↔ ∃ x, p x := Decidable.not_forall_not From 3e1e5737321bf246eb2737892f1d0728b1192b66 Mon Sep 17 00:00:00 2001 From: Xavier Roblot Date: Sat, 26 Oct 2024 11:58:43 +0000 Subject: [PATCH 442/505] feat(NumberField/Discriminant): Compute some lattices covolume (#18246) This PR is part of the proof of the Analytic Class Number Formula. --- .../NumberField/Discriminant/Basic.lean | 19 +++++++++++++++++++ 1 file changed, 19 insertions(+) diff --git a/Mathlib/NumberTheory/NumberField/Discriminant/Basic.lean b/Mathlib/NumberTheory/NumberField/Discriminant/Basic.lean index 69a827b17b1c0..013f14fce60b1 100644 --- a/Mathlib/NumberTheory/NumberField/Discriminant/Basic.lean +++ b/Mathlib/NumberTheory/NumberField/Discriminant/Basic.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Xavier Roblot. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Xavier Roblot -/ +import Mathlib.Algebra.Module.ZLattice.Covolume import Mathlib.Data.Real.Pi.Bounds import Mathlib.NumberTheory.NumberField.CanonicalEmbedding.ConvexBody import Mathlib.Tactic.Rify @@ -72,6 +73,24 @@ theorem _root_.NumberField.mixedEmbedding.volume_fundamentalDomain_latticeBasis stdBasis_repr_eq_matrixToStdBasis_mul K _ (fun _ => rfl)] rfl +theorem _root_.NumberField.mixedEmbedding.covolume_integerLattice : + ZLattice.covolume (mixedEmbedding.integerLattice K) = + (2 ⁻¹) ^ nrComplexPlaces K * √|discr K| := by + rw [ZLattice.covolume_eq_measure_fundamentalDomain _ _ (fundamentalDomain_integerLattice K), + volume_fundamentalDomain_latticeBasis, ENNReal.toReal_mul, ENNReal.toReal_pow, + ENNReal.toReal_inv, toReal_ofNat, ENNReal.coe_toReal, Real.coe_sqrt, coe_nnnorm, + Int.norm_eq_abs] + +theorem _root_.NumberField.mixedEmbedding.covolume_idealLattice (I : (FractionalIdeal (𝓞 K)⁰ K)ˣ) : + ZLattice.covolume (mixedEmbedding.idealLattice K I) = + (FractionalIdeal.absNorm (I : FractionalIdeal (𝓞 K)⁰ K)) * + (2 ⁻¹) ^ nrComplexPlaces K * √|discr K| := by + rw [ZLattice.covolume_eq_measure_fundamentalDomain _ _ (fundamentalDomain_idealLattice K I), + volume_fundamentalDomain_fractionalIdealLatticeBasis, volume_fundamentalDomain_latticeBasis, + ENNReal.toReal_mul, ENNReal.toReal_mul, ENNReal.toReal_pow, ENNReal.toReal_inv, toReal_ofNat, + ENNReal.coe_toReal, Real.coe_sqrt, coe_nnnorm, Int.norm_eq_abs, + ENNReal.toReal_ofReal (Rat.cast_nonneg.mpr (FractionalIdeal.absNorm_nonneg I.val)), mul_assoc] + theorem exists_ne_zero_mem_ideal_of_norm_le_mul_sqrt_discr (I : (FractionalIdeal (𝓞 K)⁰ K)ˣ) : ∃ a ∈ (I : FractionalIdeal (𝓞 K)⁰ K), a ≠ 0 ∧ |Algebra.norm ℚ (a : K)| ≤ FractionalIdeal.absNorm I.1 * (4 / π) ^ nrComplexPlaces K * From 8a7de7ec693d71990f0007e58c6e27f18c7290c7 Mon Sep 17 00:00:00 2001 From: Jiang Jiedong Date: Sat, 26 Oct 2024 13:22:38 +0000 Subject: [PATCH 443/505] feat(FieldTheory): define `IsConjRoot` and prove its basic properties (#17783) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We define the property `IsConjRoot K x y` as two elements having the same minimal polynomial over `K`. We show that it is equivalent to the existence of an algebra equivalence `σ : L ≃ₐ[K] L` such that `y = σ x`. We also show that if `x` is a separable element over `K` and the minimal polynomial of `x` splits in `L`, then `x` is not in the `K` iff there exists a different conjugate root of `x` in `L` over `K`. Most of the properties have a similar version already proved in Mathlib. Co-authored-by: Jiang Jiedong <107380768+jjdishere@users.noreply.github.com> --- Mathlib.lean | 1 + Mathlib/FieldTheory/Minpoly/IsConjRoot.lean | 374 ++++++++++++++++++++ 2 files changed, 375 insertions(+) create mode 100644 Mathlib/FieldTheory/Minpoly/IsConjRoot.lean diff --git a/Mathlib.lean b/Mathlib.lean index 81e1f700393e7..1c4fb5f69a188 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2799,6 +2799,7 @@ import Mathlib.FieldTheory.KummerExtension import Mathlib.FieldTheory.Laurent import Mathlib.FieldTheory.Minpoly.Basic import Mathlib.FieldTheory.Minpoly.Field +import Mathlib.FieldTheory.Minpoly.IsConjRoot import Mathlib.FieldTheory.Minpoly.IsIntegrallyClosed import Mathlib.FieldTheory.Minpoly.MinpolyDiv import Mathlib.FieldTheory.MvPolynomial diff --git a/Mathlib/FieldTheory/Minpoly/IsConjRoot.lean b/Mathlib/FieldTheory/Minpoly/IsConjRoot.lean new file mode 100644 index 0000000000000..d9c57dbc10fda --- /dev/null +++ b/Mathlib/FieldTheory/Minpoly/IsConjRoot.lean @@ -0,0 +1,374 @@ +/- +Copyright (c) 2024 Jiedong Jiang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jiedong Jiang +-/ +import Mathlib.FieldTheory.Minpoly.Basic +import Mathlib.FieldTheory.Adjoin +import Mathlib.FieldTheory.Normal + +/-! +# Conjugate roots + +Given two elements `x` and `y` of some `K`-algebra, these two elements are *conjugate roots* +over `K` if they have the same minimal polynomial over `K`. + +## Main definitions + +* `IsConjRoot`: `IsConjRoot K x y` means `y` is a conjugate root of `x` over `K`. + +## Main results + +* `isConjRoot_iff_exists_algEquiv`: Let `L / K` be a normal field extension. For any two elements +`x` and `y` in `L`, `IsConjRoot K x y` is equivalent to the existence of an algebra equivalence +`σ : L ≃ₐ[K] L` such that `y = σ x`. +* `not_mem_iff_exists_ne_and_isConjRoot`: Let `L / K` be a field extension. If `x` is a separable +element over `K` and the minimal polynomial of `x` splits in `L`, then `x` is not in the `K` iff +there exists a different conjugate root of `x` in `L` over `K`. + +## TODO +* Move `IsConjRoot` to earlier files and refactor the theorems in field theory using `IsConjRoot`. + +* Prove `IsConjRoot.smul`, if `x` and `y` are conjugate roots, then so are `r • x` and `r • y`. + +## Tags +conjugate root, minimal polynomial +-/ + + +open Polynomial minpoly IntermediateField + +variable {R K L S A B : Type*} [CommRing R] [CommRing S] [Ring A] [Ring B] [Field K] [Field L] +variable [Algebra R S] [Algebra R A] [Algebra R B] +variable [Algebra K S] [Algebra K L] [Algebra K A] [Algebra L S] + +variable (R) in +/-- +We say that `y` is a conjugate root of `x` over `K` if the minimal polynomial of `x` is the +same as the minimal polynomial of `y`. +-/ +def IsConjRoot (x y : A) : Prop := minpoly R x = minpoly R y + +/-- +The definition of conjugate roots. +-/ +theorem isConjRoot_def {x y : A} : IsConjRoot R x y ↔ minpoly R x = minpoly R y := Iff.rfl + +namespace IsConjRoot + +/-- +Every element is a conjugate root of itself. +-/ +@[refl] theorem refl {x : A} : IsConjRoot R x x := rfl + +/-- +If `y` is a conjugate root of `x`, then `x` is also a conjugate root of `y`. +-/ +@[symm] theorem symm {x y : A} (h : IsConjRoot R x y) : IsConjRoot R y x := Eq.symm h + +/-- +If `y` is a conjugate root of `x` and `z` is a conjugate root of `y`, then `z` is a conjugate +root of `x`. +-/ +@[trans] theorem trans {x y z: A} (h₁ : IsConjRoot R x y) (h₂ : IsConjRoot R y z) : + IsConjRoot R x z := Eq.trans h₁ h₂ + +variable (R A) in +/-- +The setoid structure on `A` defined by the equivalence relation of `IsConjRoot R · ·`. +-/ +def setoid : Setoid A where + r := IsConjRoot R + iseqv := ⟨fun _ => refl, symm, trans⟩ + +/-- +Let `p` be the minimal polynomial of `x`. If `y` is a conjugate root of `x`, then `p y = 0`. +-/ +theorem aeval_eq_zero {x y : A} (h : IsConjRoot R x y) : aeval y (minpoly R x) = 0 := + h ▸ minpoly.aeval R y + +/-- +Let `r` be an element of the base ring. If `y` is a conjugate root of `x`, then `y + r` is a +conjugate root of `x + r`. +-/ +theorem add_algebraMap {x y : S} (r : K) (h : IsConjRoot K x y) : + IsConjRoot K (x + algebraMap K S r) (y + algebraMap K S r) := by + rw [isConjRoot_def, minpoly.add_algebraMap x r, minpoly.add_algebraMap y r, h] + +/-- +Let `r` be an element of the base ring. If `y` is a conjugate root of `x`, then `y - r` is a +conjugate root of `x - r`. +-/ +theorem sub_algebraMap {x y : S} (r : K) (h : IsConjRoot K x y) : + IsConjRoot K (x - algebraMap K S r) (y - algebraMap K S r) := by + simpa only [sub_eq_add_neg, map_neg] using add_algebraMap (-r) h + +/-- +If `y` is a conjugate root of `x`, then `-y` is a conjugate root of `-x`. +-/ +theorem neg {x y : S} (h : IsConjRoot K x y) : + IsConjRoot K (-x) (-y) := by + rw [isConjRoot_def, minpoly.neg x, minpoly.neg y, h] + +end IsConjRoot + +open IsConjRoot + +/-- +A variant of `isConjRoot_algHom_iff`, only assuming `Function.Injective f`, +instead of `DivisionRing A`. +If `y` is a conjugate root of `x` and `f` is an injective `R`-algebra homomorphism, then `f y` is +a conjugate root of `f x`. +-/ +theorem isConjRoot_algHom_iff_of_injective {x y : A} {f : A →ₐ[R] B} + (hf : Function.Injective f) : IsConjRoot R (f x) (f y) ↔ IsConjRoot R x y := by + rw [isConjRoot_def, isConjRoot_def, algHom_eq f hf, algHom_eq f hf] + +/-- +If `y` is a conjugate root of `x` in some division ring and `f` is a `R`-algebra homomorphism, then +`f y` is a conjugate root of `f x`. +-/ +theorem isConjRoot_algHom_iff {A} [DivisionRing A] [Algebra R A] + [Nontrivial B] {x y : A} (f : A →ₐ[R] B) : IsConjRoot R (f x) (f y) ↔ IsConjRoot R x y := + isConjRoot_algHom_iff_of_injective f.injective + +/-- +Let `p` be the minimal polynomial of an integral element `x`. If `p y` = 0, then `y` is a +conjugate root of `x`. +-/ +theorem isConjRoot_of_aeval_eq_zero [IsDomain A] {x y : A} (hx : IsIntegral K x) + (h : aeval y (minpoly K x) = 0) : IsConjRoot K x y := + minpoly.eq_of_irreducible_of_monic (minpoly.irreducible hx) h (minpoly.monic hx) + +/-- +Let `p` be the minimal polynomial of an integral element `x`. Then `y` is a conjugate root of `x` +if and only if `p y = 0`. +-/ +theorem isConjRoot_iff_aeval_eq_zero [IsDomain A] {x y : A} + (h : IsIntegral K x) : IsConjRoot K x y ↔ aeval y (minpoly K x) = 0 := + ⟨IsConjRoot.aeval_eq_zero, isConjRoot_of_aeval_eq_zero h⟩ + +/-- +Let `s` be an `R`-algebra isomorphism. Then `s x` is a conjugate root of `x`. +-/ +@[simp] +theorem isConjRoot_of_algEquiv (x : A) (s : A ≃ₐ[R] A) : IsConjRoot R x (s x) := + Eq.symm (minpoly.algEquiv_eq s x) + +/-- +A variant of `isConjRoot_of_algEquiv`. +Let `s` be an `R`-algebra isomorphism. Then `x` is a conjugate root of `s x`. +-/ +@[simp] +theorem isConjRoot_of_algEquiv' (x : A) (s : A ≃ₐ[R] A) : IsConjRoot R (s x) x := + (minpoly.algEquiv_eq s x) + +/-- +Let `s₁` and `s₂` be two `R`-algebra isomorphisms. Then `s₂ x` is a conjugate root of `s₁ x`. +-/ +@[simp] +theorem isConjRoot_of_algEquiv₂ (x : A) (s₁ s₂ : A ≃ₐ[R] A) : IsConjRoot R (s₁ x) (s₂ x) := + isConjRoot_def.mpr <| (minpoly.algEquiv_eq s₂ x) ▸ (minpoly.algEquiv_eq s₁ x) + +/-- +Let `L / K` be a normal field extension. For any two elements `x` and `y` in `L`, if `y` is a +conjugate root of `x`, then there exists a `K`-automorphism `σ : L ≃ₐ[K] L` such +that `σ y = x`. +-/ +theorem IsConjRoot.exists_algEquiv [Normal K L] {x y: L} (h : IsConjRoot K x y) : + ∃ σ : L ≃ₐ[K] L, σ y = x := by + obtain ⟨σ, hσ⟩ := + exists_algHom_of_splits_of_aeval (normal_iff.mp inferInstance) (h ▸ minpoly.aeval K x) + exact ⟨AlgEquiv.ofBijective σ (σ.normal_bijective _ _ _), hσ⟩ + +/-- +Let `L / K` be a normal field extension. For any two elements `x` and `y` in `L`, `y` is a +conjugate root of `x` if and only if there exists a `K`-automorphism `σ : L ≃ₐ[K] L` such +that `σ y = x`. +-/ +theorem isConjRoot_iff_exists_algEquiv [Normal K L] {x y : L} : + IsConjRoot K x y ↔ ∃ σ : L ≃ₐ[K] L, σ y = x := + ⟨exists_algEquiv, fun ⟨_, h⟩ => h ▸ (isConjRoot_of_algEquiv _ _).symm⟩ + +/-- +Let `L / K` be a normal field extension. For any two elements `x` and `y` in `L`, `y` is a +conjugate root of `x` if and only if `x` and `y` falls in the same orbit of the action of Galois +group. +-/ +theorem isConjRoot_iff_orbitRel [Normal K L] {x y : L} : + IsConjRoot K x y ↔ MulAction.orbitRel (L ≃ₐ[K] L) L x y:= + (isConjRoot_iff_exists_algEquiv) + +variable [IsDomain S] + +/-- +Let `S / L / K` be a tower of extensions. For any two elements `y` and `x` in `S`, if `y` is a +conjugate root of `x` over `L`, then `y` is also a conjugate root of `x` over +`K`. +-/ +theorem IsConjRoot.of_isScalarTower [IsScalarTower K L S] {x y : S} (hx : IsIntegral K x) + (h : IsConjRoot L x y) : IsConjRoot K x y := + isConjRoot_of_aeval_eq_zero hx <| minpoly.aeval_of_isScalarTower K x y (aeval_eq_zero h) + +/-- +`y` is a conjugate root of `x` over `K` if and only if `y` is a root of the minimal polynomial of +`x`. This is variant of `isConjRoot_iff_aeval_eq_zero`. +-/ +theorem isConjRoot_iff_mem_minpoly_aroots {x y : S} (h : IsIntegral K x) : + IsConjRoot K x y ↔ y ∈ (minpoly K x).aroots S := by + rw [Polynomial.mem_aroots, isConjRoot_iff_aeval_eq_zero h] + simp only [iff_and_self] + exact fun _ => minpoly.ne_zero h + +/-- +`y` is a conjugate root of `x` over `K` if and only if `y` is a root of the minimal polynomial of +`x`. This is variant of `isConjRoot_iff_aeval_eq_zero`. +-/ +theorem isConjRoot_iff_mem_minpoly_rootSet {x y : S} + (h : IsIntegral K x) : IsConjRoot K x y ↔ y ∈ (minpoly K x).rootSet S := + (isConjRoot_iff_mem_minpoly_aroots h).trans (by simp [rootSet]) + +namespace IsConjRoot + +/-- +If `y` is a conjugate root of an integral element `x` over `R`, then `y` is also integral +over `R`. +-/ +theorem isIntegral {x y : A} (hx : IsIntegral R x) (h : IsConjRoot R x y) : + IsIntegral R y := + ⟨minpoly R x, minpoly.monic hx, h ▸ minpoly.aeval R y⟩ + +/-- +A variant of `IsConjRoot.eq_of_isConjRoot_algebraMap`, only assuming `Nontrivial R`, +`NoZeroSMulDivisors R A` and `Function.Injective (algebraMap R A)` instead of `Field R`. If `x` is a +conjugate root of some element `algebraMap R S r` in the image of the base ring, then +`x = algebraMap R S r`. +-/ +theorem eq_algebraMap_of_injective [Nontrivial R] [NoZeroSMulDivisors R S] {r : R} {x : S} + (h : IsConjRoot R (algebraMap R S r) x) (hf : Function.Injective (algebraMap R S)) : + x = algebraMap R S r := by + rw [IsConjRoot, minpoly.eq_X_sub_C_of_algebraMap_inj _ hf] at h + have : x ∈ (X - C r).aroots S := by + rw [mem_aroots] + simp [X_sub_C_ne_zero, h ▸ minpoly.aeval R x] + simpa [aroots_X_sub_C] using this + +/-- +If `x` is a conjugate root of some element `algebraMap R S r` in the image of the base ring, then +`x = algebraMap R S r`. +-/ +theorem eq_algebraMap {r : K} {x : S} (h : IsConjRoot K (algebraMap K S r) x) : + x = algebraMap K S r := + eq_algebraMap_of_injective h (algebraMap K S).injective + +/-- +A variant of `IsConjRoot.eq_zero`, only assuming `Nontrivial R`, +`NoZeroSMulDivisors R A` and `Function.Injective (algebraMap R A)` instead of `Field R`. If `x` is a +conjugate root of `0`, then `x = 0`. +-/ +theorem eq_zero_of_injective [Nontrivial R] [NoZeroSMulDivisors R S] {x : S} (h : IsConjRoot R 0 x) + (hf : Function.Injective (algebraMap R S)) : x = 0 := + (algebraMap R S).map_zero ▸ (eq_algebraMap_of_injective ((algebraMap R S).map_zero ▸ h) hf) + +/-- +If `x` is a conjugate root of `0`, then `x = 0`. +-/ +theorem eq_zero {x : S} (h : IsConjRoot K 0 x) : x = 0 := + eq_zero_of_injective h (algebraMap K S).injective + +end IsConjRoot + +/-- +A variant of `IsConjRoot.eq_of_isConjRoot_algebraMap`, only assuming `Nontrivial R`, +`NoZeroSMulDivisors R A` and `Function.Injective (algebraMap R A)` instead of `Field R`. If `x` is a +conjugate root of some element `algebraMap R S r` in the image of the base ring, then +`x = algebraMap R S r`. +-/ +theorem isConjRoot_iff_eq_algebraMap_of_injective [Nontrivial R] [NoZeroSMulDivisors R S] {r : R} + {x : S} (hf : Function.Injective (algebraMap R S)) : + IsConjRoot R (algebraMap R S r) x ↔ x = algebraMap R S r := + ⟨fun h => eq_algebraMap_of_injective h hf, fun h => h.symm ▸ rfl⟩ + +/-- +An element `x` is a conjugate root of some element `algebraMap R S r` in the image of the base ring +if and only if `x = algebraMap R S r`. +-/ +@[simp] +theorem isConjRoot_iff_eq_algebraMap {r : K} {x : S} : + IsConjRoot K (algebraMap K S r) x ↔ x = algebraMap K S r := + isConjRoot_iff_eq_algebraMap_of_injective (algebraMap K S).injective + +/-- +A variant of `isConjRoot_iff_eq_algebraMap`. +an element `algebraMap R S r` in the image of the base ring is a conjugate root of an element `x` +if and only if `x = algebraMap R S r`. +-/ +@[simp] +theorem isConjRoot_iff_eq_algebraMap' {r : K} {x : S} : + IsConjRoot K x (algebraMap K S r) ↔ x = algebraMap K S r := + eq_comm.trans <| isConjRoot_iff_eq_algebraMap_of_injective (algebraMap K S).injective + +/-- +A variant of `IsConjRoot.iff_eq_zero`, only assuming `Nontrivial R`, +`NoZeroSMulDivisors R A` and `Function.Injective (algebraMap R A)` instead of `Field R`. `x` is a +conjugate root of `0` if and only if `x = 0`. +-/ +theorem isConjRoot_zero_iff_eq_zero_of_injective [Nontrivial R] {x : S} [NoZeroSMulDivisors R S] + (hf : Function.Injective (algebraMap R S)) : IsConjRoot R 0 x ↔ x = 0 := + ⟨fun h => eq_zero_of_injective h hf, fun h => h.symm ▸ rfl⟩ + +/-- +`x` is a conjugate root of `0` if and only if `x = 0`. +-/ +@[simp] +theorem isConjRoot_zero_iff_eq_zero {x : S} : IsConjRoot K 0 x ↔ x = 0 := + isConjRoot_zero_iff_eq_zero_of_injective (algebraMap K S).injective + +/-- +A variant of `IsConjRoot.iff_eq_zero`. `0` is a conjugate root of `x` if and only if `x = 0`. +-/ +@[simp] +theorem isConjRoot_zero_iff_eq_zero' {x : S} : IsConjRoot K x 0 ↔ x = 0 := + eq_comm.trans <| isConjRoot_zero_iff_eq_zero_of_injective (algebraMap K S).injective + +namespace IsConjRoot + +/-- +A variant of `IsConjRoot.ne_zero`, only assuming `Nontrivial R`, +`NoZeroSMulDivisors R A` and `Function.Injective (algebraMap R A)` instead of `Field R`. If `y` is +a conjugate root of a nonzero element `x`, then `y` is not zero. +-/ +theorem ne_zero_of_injective [Nontrivial R] [NoZeroSMulDivisors R S] {x y : S} (hx : x ≠ 0) + (h : IsConjRoot R x y) (hf : Function.Injective (algebraMap R S)) : y ≠ 0 := + fun g => hx (eq_zero_of_injective (g ▸ h.symm) hf) + +/-- +If `y` is a conjugate root of a nonzero element `x`, then `y` is not zero. +-/ +theorem ne_zero {x y : S} (hx : x ≠ 0) (h : IsConjRoot K x y) : y ≠ 0 := + ne_zero_of_injective hx h (algebraMap K S).injective + +end IsConjRoot + +/-- +Let `L / K` be a field extension. If `x` is a separable element over `K` and the minimal polynomial +of `x` splits in `L`, then `x` is not in `K` if and only if there exists a conjugate +root of `x` over `K` in `L` which is not equal to `x` itself. +-/ +theorem not_mem_iff_exists_ne_and_isConjRoot {x : L} (h : IsSeparable K x) + (sp : (minpoly K x).Splits (algebraMap K L)) : + x ∉ (⊥ : Subalgebra K L) ↔ ∃ y : L, x ≠ y ∧ IsConjRoot K x y := by + calc + _ ↔ 2 ≤ (minpoly K x).natDegree := (minpoly.two_le_natDegree_iff h.isIntegral).symm + _ ↔ 2 ≤ Fintype.card ((minpoly K x).rootSet L) := + (Polynomial.card_rootSet_eq_natDegree h sp) ▸ Iff.rfl + _ ↔ Nontrivial ((minpoly K x).rootSet L) := Fintype.one_lt_card_iff_nontrivial + _ ↔ ∃ y : ((minpoly K x).rootSet L), ↑y ≠ x := + (nontrivial_iff_exists_ne ⟨x, mem_rootSet.mpr ⟨minpoly.ne_zero h.isIntegral, + minpoly.aeval K x⟩⟩).trans ⟨fun ⟨y, hy⟩ => ⟨y, Subtype.coe_ne_coe.mpr hy⟩, + fun ⟨y, hy⟩ => ⟨y, Subtype.coe_ne_coe.mp hy⟩⟩ + _ ↔ _ := + ⟨fun ⟨⟨y, hy⟩, hne⟩ => ⟨y, ⟨hne.symm, + (isConjRoot_iff_mem_minpoly_rootSet h.isIntegral).mpr hy⟩⟩, + fun ⟨y, hne, hy⟩ => ⟨⟨y, + (isConjRoot_iff_mem_minpoly_rootSet h.isIntegral).mp hy⟩, hne.symm⟩⟩ From a9b268f6d9ab29307631902be31bb9aca0ddf72e Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Sat, 26 Oct 2024 13:22:39 +0000 Subject: [PATCH 444/505] =?UTF-8?q?refactor:=20change=20the=20definition?= =?UTF-8?q?=20of=20`1=20:=20Submodule=20R=20A`=20to=20the=20range=20of=20`?= =?UTF-8?q?(=C2=B7=20=E2=80=A2=201)`=20(#18092)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Previously, it's defined to be the range of `algebraMap R A`. In the noncommutative setting, we want to write `1 : Ideal R` where `Ideal R := Submodule R R`, but `R` isn't an `R`-algebra if `R` is not commutative. If we were to introduce a new typeclass by removing the `commutes'` field from `Algebra`, we could keep the current definition, but I'd argue that it's still better to use SMul to define `1 : Submodule R A`, because `Submodule` only depends on the Module/SMul, not on the algebraMap/RingHom. Requires fixes in 11 other files. --- Mathlib/Algebra/Algebra/Bilinear.lean | 3 + Mathlib/Algebra/Algebra/Operations.lean | 75 +++++++++++-------- Mathlib/Algebra/Algebra/Subalgebra/Basic.lean | 7 +- .../BilinearForm/DualLattice.lean | 8 +- .../CliffordAlgebra/Grading.lean | 4 +- .../Dimension/FreeAndStrongRankCondition.lean | 3 +- .../LinearAlgebra/ExteriorAlgebra/Basic.lean | 3 +- Mathlib/LinearAlgebra/LinearDisjoint.lean | 4 +- .../LinearAlgebra/TensorAlgebra/Basic.lean | 2 +- .../TensorProduct/Submodule.lean | 20 +++-- .../RingTheory/DedekindDomain/Different.lean | 14 ++-- .../DedekindDomain/IntegralClosure.lean | 2 +- Mathlib/RingTheory/FractionalIdeal/Basic.lean | 2 +- 13 files changed, 87 insertions(+), 60 deletions(-) diff --git a/Mathlib/Algebra/Algebra/Bilinear.lean b/Mathlib/Algebra/Algebra/Bilinear.lean index 65e3fda3f286f..82c8c45e63e1c 100644 --- a/Mathlib/Algebra/Algebra/Bilinear.lean +++ b/Mathlib/Algebra/Algebra/Bilinear.lean @@ -215,6 +215,9 @@ theorem pow_mulRight (a : A) (n : ℕ) : mulRight R a ^ n = mulRight R (a ^ n) : exact LinearMap.coe_injective (((mulRight R a).coe_pow n).symm ▸ mul_right_iterate a n) +theorem toSpanSingleton_eq_algebra_linearMap : toSpanSingleton R A 1 = Algebra.linearMap R A := by + ext; simp + end Semiring section Ring diff --git a/Mathlib/Algebra/Algebra/Operations.lean b/Mathlib/Algebra/Algebra/Operations.lean index ce82654752d9f..a962fa9ac33d8 100644 --- a/Mathlib/Algebra/Algebra/Operations.lean +++ b/Mathlib/Algebra/Algebra/Operations.lean @@ -36,6 +36,11 @@ It is proved that `Submodule R A` is a semiring, and also an algebra over `Set A Additionally, in the `Pointwise` locale we promote `Submodule.pointwiseDistribMulAction` to a `MulSemiringAction` as `Submodule.pointwiseMulSemiringAction`. +When `R` is not necessarily commutative, and `A` is merely a `R`-module with a ring structure +such that `IsScalarTower R A A` holds (equivalent to the data of a ring homomorphism `R →+* A` +by `ringHomEquivModuleIsScalarTower`), we can still define `1 : Submodule R A` and +`Mul (Submodule R A)`, but `1` is only a left identity, not necessarily a right one. + ## Tags multiplication of submodules, division of submodules, submodule semiring @@ -62,49 +67,57 @@ end SubMulAction namespace Submodule -variable {ι : Sort uι} -variable {R : Type u} [CommSemiring R] - -section Ring +section Module -variable {A : Type v} [Semiring A] [Algebra R A] -variable (S T : Set A) {M N P Q : Submodule R A} {m n : A} +variable {R : Type u} [Semiring R] {A : Type v} [Semiring A] [Module R A] -/-- `1 : Submodule R A` is the submodule R of A. -/ +/-- `1 : Submodule R A` is the submodule `R ∙ 1` of A. +TODO: potentially change this back to `LinearMap.range (Algebra.linearMap R A)` +once a version of `Algebra` without the `commutes'` field is introduced. +See issue #18110. +-/ instance one : One (Submodule R A) := - -- Porting note: `f.range` notation doesn't work - ⟨LinearMap.range (Algebra.linearMap R A)⟩ + ⟨LinearMap.range (LinearMap.toSpanSingleton R A 1)⟩ -theorem one_eq_range : (1 : Submodule R A) = LinearMap.range (Algebra.linearMap R A) := - rfl +theorem one_eq_span : (1 : Submodule R A) = R ∙ 1 := + (LinearMap.span_singleton_eq_range _ _ _).symm theorem le_one_toAddSubmonoid : 1 ≤ (1 : Submodule R A).toAddSubmonoid := by rintro x ⟨n, rfl⟩ - exact ⟨n, map_natCast (algebraMap R A) n⟩ - -theorem algebraMap_mem (r : R) : algebraMap R A r ∈ (1 : Submodule R A) := - LinearMap.mem_range_self (Algebra.linearMap R A) _ - -@[simp] -theorem mem_one {x : A} : x ∈ (1 : Submodule R A) ↔ ∃ y, algebraMap R A y = x := - Iff.rfl + exact ⟨n, show (n : R) • (1 : A) = n by rw [Nat.cast_smul_eq_nsmul, nsmul_one]⟩ @[simp] theorem toSubMulAction_one : (1 : Submodule R A).toSubMulAction = 1 := - SetLike.ext fun _ => mem_one.trans SubMulAction.mem_one'.symm - -theorem one_eq_span : (1 : Submodule R A) = R ∙ 1 := by - apply Submodule.ext - intro a - simp only [mem_one, mem_span_singleton, Algebra.smul_def, mul_one] + SetLike.ext fun _ ↦ by rw [one_eq_span, SubMulAction.mem_one]; exact mem_span_singleton theorem one_eq_span_one_set : (1 : Submodule R A) = span R 1 := one_eq_span -theorem one_le : (1 : Submodule R A) ≤ P ↔ (1 : A) ∈ P := by +theorem one_le {P : Submodule R A} : (1 : Submodule R A) ≤ P ↔ (1 : A) ∈ P := by -- Porting note: simpa no longer closes refl goals, so added `SetLike.mem_coe` simp only [one_eq_span, span_le, Set.singleton_subset_iff, SetLike.mem_coe] +end Module + +variable {ι : Sort uι} +variable {R : Type u} [CommSemiring R] + +section AlgebraSemiring + +variable {A : Type v} [Semiring A] [Algebra R A] +variable (S T : Set A) {M N P Q : Submodule R A} {m n : A} + +theorem one_eq_range : (1 : Submodule R A) = LinearMap.range (Algebra.linearMap R A) := by + rw [one_eq_span, LinearMap.span_singleton_eq_range, + LinearMap.toSpanSingleton_eq_algebra_linearMap] + +theorem algebraMap_mem (r : R) : algebraMap R A r ∈ (1 : Submodule R A) := by + rw [one_eq_range]; exact LinearMap.mem_range_self _ _ + +@[simp] +theorem mem_one {x : A} : x ∈ (1 : Submodule R A) ↔ ∃ y, algebraMap R A y = x := by + rw [one_eq_range]; rfl + protected theorem map_one {A'} [Semiring A'] [Algebra R A'] (f : A →ₐ[R] A') : map f.toLinearMap (1 : Submodule R A) = 1 := by ext @@ -423,7 +436,7 @@ protected theorem pow_induction_on_left' {C : ∀ (n : ℕ) (x), x ∈ M ^ n → induction n generalizing x with | zero => rw [pow_zero] at hx - obtain ⟨r, rfl⟩ := hx + obtain ⟨r, rfl⟩ := mem_one.mp hx exact algebraMap r | succ n n_ih => revert hx @@ -444,7 +457,7 @@ protected theorem pow_induction_on_right' {C : ∀ (n : ℕ) (x), x ∈ M ^ n induction n generalizing x with | zero => rw [pow_zero] at hx - obtain ⟨r, rfl⟩ := hx + obtain ⟨r, rfl⟩ := mem_one.mp hx exact algebraMap r | succ n n_ih => revert hx @@ -549,9 +562,9 @@ scoped[Pointwise] attribute [instance] Submodule.pointwiseMulSemiringAction end -end Ring +end AlgebraSemiring -section CommRing +section AlgebraCommSemiring variable {A : Type v} [CommSemiring A] [Algebra R A] variable {M N : Submodule R A} {m n : A} @@ -687,6 +700,6 @@ protected theorem map_div {B : Type*} [CommSemiring B] [Algebra R B] (I J : Subm end Quotient -end CommRing +end AlgebraCommSemiring end Submodule diff --git a/Mathlib/Algebra/Algebra/Subalgebra/Basic.lean b/Mathlib/Algebra/Algebra/Subalgebra/Basic.lean index a34f329d35336..7683b0c0af36d 100644 --- a/Mathlib/Algebra/Algebra/Subalgebra/Basic.lean +++ b/Mathlib/Algebra/Algebra/Subalgebra/Basic.lean @@ -776,7 +776,9 @@ instance : Inhabited (Subalgebra R A) := ⟨⊥⟩ theorem mem_bot {x : A} : x ∈ (⊥ : Subalgebra R A) ↔ x ∈ Set.range (algebraMap R A) := Iff.rfl -theorem toSubmodule_bot : Subalgebra.toSubmodule (⊥ : Subalgebra R A) = 1 := rfl +/-- TODO: change proof to `rfl` when fixing #18110. -/ +theorem toSubmodule_bot : Subalgebra.toSubmodule (⊥ : Subalgebra R A) = 1 := + Submodule.one_eq_range.symm @[simp] theorem coe_bot : ((⊥ : Subalgebra R A) : Set A) = Set.range (algebraMap R A) := rfl @@ -799,7 +801,8 @@ theorem map_top (f : A →ₐ[R] B) : (⊤ : Subalgebra R A).map f = f.range := @[simp] theorem map_bot (f : A →ₐ[R] B) : (⊥ : Subalgebra R A).map f = ⊥ := - Subalgebra.toSubmodule_injective <| Submodule.map_one _ + Subalgebra.toSubmodule_injective <| by + simpa only [Subalgebra.map_toSubmodule, toSubmodule_bot] using Submodule.map_one _ @[simp] theorem comap_top (f : A →ₐ[R] B) : (⊤ : Subalgebra R B).comap f = ⊤ := diff --git a/Mathlib/LinearAlgebra/BilinearForm/DualLattice.lean b/Mathlib/LinearAlgebra/BilinearForm/DualLattice.lean index 24c02406a750c..8fba55a7ccff6 100644 --- a/Mathlib/LinearAlgebra/BilinearForm/DualLattice.lean +++ b/Mathlib/LinearAlgebra/BilinearForm/DualLattice.lean @@ -51,12 +51,12 @@ lemma le_flip_dualSubmodule {N₁ N₂ : Submodule R M} : This is bundled as a bilinear map in `BilinForm.dualSubmoduleToDual`. -/ noncomputable def dualSubmoduleParing {N : Submodule R M} (x : B.dualSubmodule N) (y : N) : R := - (x.prop y y.prop).choose + (Submodule.mem_one.mp <| x.prop y y.prop).choose @[simp] lemma dualSubmoduleParing_spec {N : Submodule R M} (x : B.dualSubmodule N) (y : N) : algebraMap R S (B.dualSubmoduleParing x y) = B x y := - (x.prop y y.prop).choose_spec + (Submodule.mem_one.mp <| x.prop y y.prop).choose_spec /-- The natural paring of `B.dualSubmodule N` and `N`. -/ -- TODO: Show that this is perfect when `N` is a lattice and `B` is nondegenerate. @@ -94,7 +94,7 @@ lemma dualSubmodule_span_of_basis {ι} [Finite ι] [DecidableEq ι] rw [← (B.dualBasis hB b).sum_repr x] apply sum_mem rintro i - - obtain ⟨r, hr⟩ := hx (b i) (Submodule.subset_span ⟨_, rfl⟩) + obtain ⟨r, hr⟩ := Submodule.mem_one.mp <| hx (b i) (Submodule.subset_span ⟨_, rfl⟩) simp only [dualBasis_repr_apply, ← hr, Algebra.linearMap_apply, algebraMap_smul] apply Submodule.smul_mem exact Submodule.subset_span ⟨_, rfl⟩ @@ -107,7 +107,7 @@ lemma dualSubmodule_span_of_basis {ι} [Finite ι] [DecidableEq ι] rw [← IsScalarTower.algebraMap_smul S (f j), map_smul] simp_rw [apply_dualBasis_left] rw [smul_eq_mul, mul_ite, mul_one, mul_zero, ← (algebraMap R S).map_zero, ← apply_ite] - exact ⟨_, rfl⟩ + exact Submodule.mem_one.mpr ⟨_, rfl⟩ lemma dualSubmodule_dualSubmodule_flip_of_basis {ι : Type*} [Finite ι] (hB : B.Nondegenerate) (b : Basis ι S M) : diff --git a/Mathlib/LinearAlgebra/CliffordAlgebra/Grading.lean b/Mathlib/LinearAlgebra/CliffordAlgebra/Grading.lean index 145bb1e7b4a15..e0cba62521842 100644 --- a/Mathlib/LinearAlgebra/CliffordAlgebra/Grading.lean +++ b/Mathlib/LinearAlgebra/CliffordAlgebra/Grading.lean @@ -197,8 +197,8 @@ theorem even_induction {motive : ∀ x, x ∈ evenOdd Q 0 → Prop} motive (ι Q m₁ * ι Q m₂ * x) (zero_add (0 : ZMod 2) ▸ SetLike.mul_mem_graded (ι_mul_ι_mem_evenOdd_zero Q m₁ m₂) hx)) (x : CliffordAlgebra Q) (hx : x ∈ evenOdd Q 0) : motive x hx := by - refine evenOdd_induction _ _ (motive := motive) (fun rx => ?_) add ι_mul_ι_mul x hx - rintro ⟨r, rfl⟩ + refine evenOdd_induction _ _ (motive := motive) (fun rx h => ?_) add ι_mul_ι_mul x hx + obtain ⟨r, rfl⟩ := Submodule.mem_one.mp h exact algebraMap r /-- To show a property is true on the odd parts, it suffices to show it is true on the diff --git a/Mathlib/LinearAlgebra/Dimension/FreeAndStrongRankCondition.lean b/Mathlib/LinearAlgebra/Dimension/FreeAndStrongRankCondition.lean index 43e1a1eedf450..212770aa65383 100644 --- a/Mathlib/LinearAlgebra/Dimension/FreeAndStrongRankCondition.lean +++ b/Mathlib/LinearAlgebra/Dimension/FreeAndStrongRankCondition.lean @@ -271,7 +271,8 @@ theorem rank_eq_one_iff [Nontrivial E] [Module.Free F S] : Module.rank F S = 1 obtain ⟨κ, b⟩ := Module.Free.exists_basis (R := F) (M := (⊥ : Subalgebra F E)) refine le_antisymm ?_ ?_ · have := lift_rank_range_le (Algebra.linearMap F E) - rwa [← one_eq_range, rank_self, lift_one, lift_le_one_iff] at this + rwa [← one_eq_range, rank_self, lift_one, lift_le_one_iff, + ← Algebra.toSubmodule_bot, rank_toSubmodule] at this · by_contra H rw [not_le, lt_one_iff_zero] at H haveI := mk_eq_zero_iff.1 (H ▸ b.mk_eq_rank'') diff --git a/Mathlib/LinearAlgebra/ExteriorAlgebra/Basic.lean b/Mathlib/LinearAlgebra/ExteriorAlgebra/Basic.lean index 75b12c7d91e55..b81a28f1405ee 100644 --- a/Mathlib/LinearAlgebra/ExteriorAlgebra/Basic.lean +++ b/Mathlib/LinearAlgebra/ExteriorAlgebra/Basic.lean @@ -239,7 +239,8 @@ theorem ι_range_disjoint_one : Disjoint (LinearMap.range (ι R : M →ₗ[R] ExteriorAlgebra R M)) (1 : Submodule R (ExteriorAlgebra R M)) := by rw [Submodule.disjoint_def] - rintro _ ⟨x, hx⟩ ⟨r, rfl : algebraMap R (ExteriorAlgebra R M) r = _⟩ + rintro _ ⟨x, hx⟩ h + obtain ⟨r, rfl : algebraMap R (ExteriorAlgebra R M) r = _⟩ := Submodule.mem_one.mp h rw [ι_eq_algebraMap_iff x] at hx rw [hx.2, RingHom.map_zero] diff --git a/Mathlib/LinearAlgebra/LinearDisjoint.lean b/Mathlib/LinearAlgebra/LinearDisjoint.lean index f290e4e777d86..efcf71db426f2 100644 --- a/Mathlib/LinearAlgebra/LinearDisjoint.lean +++ b/Mathlib/LinearAlgebra/LinearDisjoint.lean @@ -232,12 +232,12 @@ theorem bot_right : M.LinearDisjoint (⊥ : Submodule R S) := /-- The image of `R` in `S` is linearly disjoint with any other submodules. -/ theorem one_left : (1 : Submodule R S).LinearDisjoint N := by - rw [linearDisjoint_iff, mulMap_one_left_eq] + rw [linearDisjoint_iff, ← Algebra.toSubmodule_bot, mulMap_one_left_eq] exact N.injective_subtype.comp N.lTensorOne.injective /-- The image of `R` in `S` is linearly disjoint with any other submodules. -/ theorem one_right : M.LinearDisjoint (1 : Submodule R S) := by - rw [linearDisjoint_iff, mulMap_one_right_eq] + rw [linearDisjoint_iff, ← Algebra.toSubmodule_bot, mulMap_one_right_eq] exact M.injective_subtype.comp M.rTensorOne.injective /-- If for any finitely generated submodules `M'` of `M`, `M'` and `N` are linearly disjoint, diff --git a/Mathlib/LinearAlgebra/TensorAlgebra/Basic.lean b/Mathlib/LinearAlgebra/TensorAlgebra/Basic.lean index 9355d09c2992a..f24b144938533 100644 --- a/Mathlib/LinearAlgebra/TensorAlgebra/Basic.lean +++ b/Mathlib/LinearAlgebra/TensorAlgebra/Basic.lean @@ -288,7 +288,7 @@ theorem ι_ne_one [Nontrivial R] (x : M) : ι R x ≠ 1 := by theorem ι_range_disjoint_one : Disjoint (LinearMap.range (ι R : M →ₗ[R] TensorAlgebra R M)) (1 : Submodule R (TensorAlgebra R M)) := by - rw [Submodule.disjoint_def] + rw [Submodule.disjoint_def, Submodule.one_eq_range] rintro _ ⟨x, hx⟩ ⟨r, rfl⟩ rw [Algebra.linearMap_apply, ι_eq_algebraMap_iff] at hx rw [hx.2, map_zero] diff --git a/Mathlib/LinearAlgebra/TensorProduct/Submodule.lean b/Mathlib/LinearAlgebra/TensorProduct/Submodule.lean index 838f15025ec00..074c5389f0f01 100644 --- a/Mathlib/LinearAlgebra/TensorProduct/Submodule.lean +++ b/Mathlib/LinearAlgebra/TensorProduct/Submodule.lean @@ -121,8 +121,9 @@ theorem mulMap'_surjective : Function.Surjective (mulMap' M N) := by `i(R) ⊗[R] N →ₗ[R] N` induced by multiplication in `S`, here `i : R → S` is the structure map. This is promoted to an isomorphism of `R`-modules as `Submodule.lTensorOne`. Use that instead. -/ def lTensorOne' : (⊥ : Subalgebra R S) ⊗[R] N →ₗ[R] N := - show (1 : Submodule R S) ⊗[R] N →ₗ[R] N from - (LinearEquiv.ofEq _ _ (by rw [mulMap_range, one_mul])).toLinearMap ∘ₗ (mulMap _ N).rangeRestrict + show Subalgebra.toSubmodule ⊥ ⊗[R] N →ₗ[R] N from + (LinearEquiv.ofEq _ _ (by rw [Algebra.toSubmodule_bot, mulMap_range, one_mul])).toLinearMap ∘ₗ + (mulMap _ N).rangeRestrict variable {N} in @[simp] @@ -130,7 +131,7 @@ theorem lTensorOne'_tmul (y : R) (n : N) : N.lTensorOne' (algebraMap R _ y ⊗ₜ[R] n) = y • n := Subtype.val_injective <| by simp_rw [lTensorOne', LinearMap.coe_comp, LinearEquiv.coe_coe, Function.comp_apply, LinearEquiv.coe_ofEq_apply, LinearMap.codRestrict_apply, SetLike.val_smul, Algebra.smul_def] - exact mulMap_tmul 1 N _ _ + exact mulMap_tmul _ N _ _ variable {N} in @[simp] @@ -163,15 +164,17 @@ variable {N} in @[simp] theorem lTensorOne_symm_apply (n : N) : N.lTensorOne.symm n = 1 ⊗ₜ[R] n := rfl -theorem mulMap_one_left_eq : mulMap 1 N = N.subtype ∘ₗ N.lTensorOne.toLinearMap := +theorem mulMap_one_left_eq : + mulMap (Subalgebra.toSubmodule ⊥) N = N.subtype ∘ₗ N.lTensorOne.toLinearMap := TensorProduct.ext' fun _ _ ↦ rfl /-- If `M` is a submodule in an algebra `S` over `R`, there is the natural `R`-linear map `M ⊗[R] i(R) →ₗ[R] M` induced by multiplication in `S`, here `i : R → S` is the structure map. This is promoted to an isomorphism of `R`-modules as `Submodule.rTensorOne`. Use that instead. -/ def rTensorOne' : M ⊗[R] (⊥ : Subalgebra R S) →ₗ[R] M := - show M ⊗[R] (1 : Submodule R S) →ₗ[R] M from - (LinearEquiv.ofEq _ _ (by rw [mulMap_range, mul_one])).toLinearMap ∘ₗ (mulMap M _).rangeRestrict + show M ⊗[R] Subalgebra.toSubmodule ⊥ →ₗ[R] M from + (LinearEquiv.ofEq _ _ (by rw [Algebra.toSubmodule_bot, mulMap_range, mul_one])).toLinearMap ∘ₗ + (mulMap M _).rangeRestrict variable {M} in @[simp] @@ -180,7 +183,7 @@ theorem rTensorOne'_tmul (y : R) (m : M) : simp_rw [rTensorOne', LinearMap.coe_comp, LinearEquiv.coe_coe, Function.comp_apply, LinearEquiv.coe_ofEq_apply, LinearMap.codRestrict_apply, SetLike.val_smul] rw [Algebra.smul_def, Algebra.commutes] - exact mulMap_tmul M 1 _ _ + exact mulMap_tmul M _ _ _ variable {M} in @[simp] @@ -213,7 +216,8 @@ variable {M} in @[simp] theorem rTensorOne_symm_apply (m : M) : M.rTensorOne.symm m = m ⊗ₜ[R] 1 := rfl -theorem mulMap_one_right_eq : mulMap M 1 = M.subtype ∘ₗ M.rTensorOne.toLinearMap := +theorem mulMap_one_right_eq : + mulMap M (Subalgebra.toSubmodule ⊥) = M.subtype ∘ₗ M.rTensorOne.toLinearMap := TensorProduct.ext' fun _ _ ↦ rfl @[simp] diff --git a/Mathlib/RingTheory/DedekindDomain/Different.lean b/Mathlib/RingTheory/DedekindDomain/Different.lean index a9e9ad8924c5b..0abb74d5ac750 100644 --- a/Mathlib/RingTheory/DedekindDomain/Different.lean +++ b/Mathlib/RingTheory/DedekindDomain/Different.lean @@ -59,7 +59,7 @@ namespace Submodule lemma mem_traceDual {I : Submodule B L} {x} : x ∈ Iᵛ ↔ ∀ a ∈ I, traceForm K L x a ∈ (algebraMap A K).range := - Iff.rfl + forall₂_congr fun _ _ ↦ mem_one lemma le_traceDual_iff_map_le_one {I J : Submodule B L} : I ≤ Jᵛ ↔ ((I * J : Submodule B L).restrictScalars A).map @@ -149,14 +149,15 @@ variable [IsIntegrallyClosed A] lemma Submodule.mem_traceDual_iff_isIntegral {I : Submodule B L} {x} : x ∈ Iᵛ ↔ ∀ a ∈ I, IsIntegral A (traceForm K L x a) := - forall₂_congr (fun _ _ ↦ IsIntegrallyClosed.isIntegral_iff.symm) + forall₂_congr fun _ _ ↦ mem_one.trans IsIntegrallyClosed.isIntegral_iff.symm variable [FiniteDimensional K L] [IsIntegralClosure B A L] lemma Submodule.one_le_traceDual_one : (1 : Submodule B L) ≤ 1ᵛ := by - rw [le_traceDual_iff_map_le_one, mul_one] + rw [le_traceDual_iff_map_le_one, mul_one, one_eq_range] rintro _ ⟨x, ⟨x, rfl⟩, rfl⟩ + rw [mem_one] apply IsIntegrallyClosed.isIntegral_iff.mp apply isIntegral_trace rw [IsIntegralClosure.isIntegral_iff (A := B)] @@ -262,7 +263,7 @@ variable {A K L B} lemma mem_dual (hI : I ≠ 0) {x} : x ∈ dual A K I ↔ ∀ a ∈ I, traceForm K L x a ∈ (algebraMap A K).range := by - rw [dual, dif_neg hI]; rfl + rw [dual, dif_neg hI]; exact forall₂_congr fun _ _ ↦ mem_one variable (A K) @@ -300,6 +301,7 @@ lemma le_dual_inv_aux (hI : I ≠ 0) (hIJ : I * J ≤ 1) : J ≤ dual A K I := by rw [dual, dif_neg hI] intro x hx y hy + rw [mem_one] apply IsIntegrallyClosed.isIntegral_iff.mp apply isIntegral_trace rw [IsIntegralClosure.isIntegral_iff (A := B)] @@ -412,7 +414,7 @@ lemma coeSubmodule_differentIdeal_fractionRing simp only [← one_div, FractionalIdeal.val_eq_coe] at this rw [FractionalIdeal.coe_div (FractionalIdeal.dual_ne_zero _ _ _), FractionalIdeal.coe_dual] at this - · simpa only [FractionalIdeal.coe_one] using this + · simpa only [FractionalIdeal.coe_one, Submodule.one_eq_range] using this · exact one_ne_zero · exact one_ne_zero @@ -559,7 +561,7 @@ lemma conductor_mul_differentIdeal [NoZeroSMulDivisors A B] derivative_map, aeval_map_algebraMap, aeval_algebraMap_apply, mul_assoc, FractionalIdeal.mem_one_iff, forall_exists_index, forall_apply_eq_imp_iff] simp_rw [← IsScalarTower.toAlgHom_apply A B L x, ← AlgHom.map_adjoin_singleton] - simp only [Subalgebra.mem_map, IsScalarTower.coe_toAlgHom', + simp only [Subalgebra.mem_map, IsScalarTower.coe_toAlgHom', Submodule.one_eq_range, forall_exists_index, and_imp, forall_apply_eq_imp_iff₂, ← _root_.map_mul] exact ⟨fun H b ↦ (mul_one b) ▸ H b 1 (one_mem _), fun H _ _ _ ↦ H _⟩ diff --git a/Mathlib/RingTheory/DedekindDomain/IntegralClosure.lean b/Mathlib/RingTheory/DedekindDomain/IntegralClosure.lean index 86796148fbcf3..81f78b774ad8d 100644 --- a/Mathlib/RingTheory/DedekindDomain/IntegralClosure.lean +++ b/Mathlib/RingTheory/DedekindDomain/IntegralClosure.lean @@ -95,7 +95,7 @@ theorem IsIntegralClosure.range_le_span_dualBasis [Algebra.IsSeparable K L] {ι rintro _ ⟨i, rfl⟩ _ ⟨y, rfl⟩ simp only [LinearMap.coe_restrictScalars, linearMap_apply, LinearMap.BilinForm.flip_apply, traceForm_apply] - refine IsIntegrallyClosed.isIntegral_iff.mp ?_ + refine Submodule.mem_one.mpr <| IsIntegrallyClosed.isIntegral_iff.mp ?_ exact isIntegral_trace ((IsIntegralClosure.isIntegral A L y).algebraMap.mul (hb_int i)) theorem integralClosure_le_span_dualBasis [Algebra.IsSeparable K L] {ι : Type*} [Fintype ι] diff --git a/Mathlib/RingTheory/FractionalIdeal/Basic.lean b/Mathlib/RingTheory/FractionalIdeal/Basic.lean index 4ebe1fa201f1f..4056634b451c7 100644 --- a/Mathlib/RingTheory/FractionalIdeal/Basic.lean +++ b/Mathlib/RingTheory/FractionalIdeal/Basic.lean @@ -215,7 +215,7 @@ theorem isFractional_of_le_one (I : Submodule R P) (h : I ≤ 1) : IsFractional use 1, S.one_mem intro b hb rw [one_smul] - obtain ⟨b', b'_mem, rfl⟩ := h hb + obtain ⟨b', b'_mem, rfl⟩ := mem_one.mp (h hb) exact Set.mem_range_self b' theorem isFractional_of_le {I : Submodule R P} {J : FractionalIdeal S P} (hIJ : I ≤ J) : From b76f29f52d5528808a5e945b510d33ebbe029e88 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Sat, 26 Oct 2024 13:22:40 +0000 Subject: [PATCH 445/505] chore(RingTheory): move `Ideal.exists_pow_le_of_le_radical_of_fG` (#18189) renamed to `Ideal.exists_pow_le_of_le_radical_of_fg_radical`, and added `Ideal.exists_pow_le_of_le_radical_of_fg` Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- Mathlib/Algebra/Homology/LocalCohomology.lean | 15 +--------- Mathlib/RingTheory/Finiteness.lean | 28 +++++++++++++++++++ Mathlib/RingTheory/Ideal/Operations.lean | 14 ++++++++++ 3 files changed, 43 insertions(+), 14 deletions(-) diff --git a/Mathlib/Algebra/Homology/LocalCohomology.lean b/Mathlib/Algebra/Homology/LocalCohomology.lean index 4ea0bfce99d36..f8d4b17460856 100644 --- a/Mathlib/Algebra/Homology/LocalCohomology.lean +++ b/Mathlib/Algebra/Homology/LocalCohomology.lean @@ -206,26 +206,13 @@ def idealPowersToSelfLERadical (J : Ideal R) : ℕᵒᵖ ⥤ SelfLERadical J := variable {I J K : Ideal R} -/-- The lemma below essentially says that `idealPowersToSelfLERadical I` is initial in -`selfLERadicalDiagram I`. - -Porting note: This lemma should probably be moved to `Mathlib/RingTheory/Finiteness` -to be near `Ideal.exists_radical_pow_le_of_fg`, which it generalizes. -/ -theorem Ideal.exists_pow_le_of_le_radical_of_fG (hIJ : I ≤ J.radical) (hJ : J.radical.FG) : - ∃ k : ℕ, I ^ k ≤ J := by - obtain ⟨k, hk⟩ := J.exists_radical_pow_le_of_fg hJ - use k - calc - I ^ k ≤ J.radical ^ k := Ideal.pow_right_mono hIJ _ - _ ≤ J := hk - /-- The diagram of powers of `J` is initial in the diagram of all ideals with radical containing `J`. This uses noetherianness. -/ instance ideal_powers_initial [hR : IsNoetherian R R] : Functor.Initial (idealPowersToSelfLERadical J) where out J' := by apply (config := {allowSynthFailures := true }) zigzag_isConnected - · obtain ⟨k, hk⟩ := Ideal.exists_pow_le_of_le_radical_of_fG J'.2 (isNoetherian_def.mp hR _) + · obtain ⟨k, hk⟩ := Ideal.exists_pow_le_of_le_radical_of_fg J'.2 (isNoetherian_def.mp hR _) exact ⟨CostructuredArrow.mk (⟨⟨hk⟩⟩ : (idealPowersToSelfLERadical J).obj (op k) ⟶ J')⟩ · intro j1 j2 apply Relation.ReflTransGen.single diff --git a/Mathlib/RingTheory/Finiteness.lean b/Mathlib/RingTheory/Finiteness.lean index caa9661139eaf..fad4868eda18b 100644 --- a/Mathlib/RingTheory/Finiteness.lean +++ b/Mathlib/RingTheory/Finiteness.lean @@ -493,6 +493,34 @@ theorem exists_radical_pow_le_of_fg {R : Type*} [CommSemiring R] (I : Ideal R) ( rw [add_comm, Nat.add_sub_assoc h.le] apply Nat.le_add_right +theorem exists_pow_le_of_le_radical_of_fg_radical {R : Type*} [CommSemiring R] {I J : Ideal R} + (hIJ : I ≤ J.radical) (hJ : J.radical.FG) : + ∃ k : ℕ, I ^ k ≤ J := by + obtain ⟨k, hk⟩ := J.exists_radical_pow_le_of_fg hJ + use k + calc + I ^ k ≤ J.radical ^ k := Ideal.pow_right_mono hIJ _ + _ ≤ J := hk + +@[deprecated (since := "2024-10-24")] +alias exists_pow_le_of_le_radical_of_fG := exists_pow_le_of_le_radical_of_fg_radical + +lemma exists_pow_le_of_le_radical_of_fg {R : Type*} [CommSemiring R] {I J : Ideal R} + (h' : I ≤ J.radical) (h : I.FG) : + ∃ n : ℕ, I ^ n ≤ J := by + revert h' + apply Submodule.fg_induction _ _ _ _ _ I h + · intro x hJ + simp only [Ideal.submodule_span_eq, Ideal.span_le, + Set.singleton_subset_iff, SetLike.mem_coe] at hJ + obtain ⟨n, hn⟩ := hJ + refine ⟨n, by simpa [Ideal.span_singleton_pow, Ideal.span_le]⟩ + · intros I₁ I₂ h₁ h₂ hJ + obtain ⟨n₁, hn₁⟩ := h₁ (le_sup_left.trans hJ) + obtain ⟨n₂, hn₂⟩ := h₂ (le_sup_right.trans hJ) + use n₁ + n₂ + exact Ideal.sup_pow_add_le_pow_sup_pow.trans (sup_le hn₁ hn₂) + end Ideal section ModuleAndAlgebra diff --git a/Mathlib/RingTheory/Ideal/Operations.lean b/Mathlib/RingTheory/Ideal/Operations.lean index 8a168c723e1d3..176f29b389039 100644 --- a/Mathlib/RingTheory/Ideal/Operations.lean +++ b/Mathlib/RingTheory/Ideal/Operations.lean @@ -603,6 +603,20 @@ theorem pow_right_mono {I J : Ideal R} (e : I ≤ J) (n : ℕ) : I ^ n ≤ J ^ n · rw [pow_succ, pow_succ] exact Ideal.mul_mono hn e +lemma sup_pow_add_le_pow_sup_pow {R} [CommSemiring R] {I J : Ideal R} {n m : ℕ} : + (I ⊔ J) ^ (n + m) ≤ I ^ n ⊔ J ^ m := by + rw [← Ideal.add_eq_sup, ← Ideal.add_eq_sup, add_pow, Ideal.sum_eq_sup] + apply Finset.sup_le + intros i hi + by_cases hn : n ≤ i + · exact (Ideal.mul_le_right.trans (Ideal.mul_le_right.trans + ((Ideal.pow_le_pow_right hn).trans le_sup_left))) + · refine (Ideal.mul_le_right.trans (Ideal.mul_le_left.trans + ((Ideal.pow_le_pow_right ?_).trans le_sup_right))) + simp only [Finset.mem_range, Nat.lt_succ] at hi + rw [Nat.le_sub_iff_add_le hi] + nlinarith + @[simp] theorem mul_eq_bot {R : Type*} [CommSemiring R] [NoZeroDivisors R] {I J : Ideal R} : I * J = ⊥ ↔ I = ⊥ ∨ J = ⊥ := From e2338ffc272658ed73b8c5a708aafc386b30e1b7 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Sat, 26 Oct 2024 13:47:49 +0000 Subject: [PATCH 446/505] feat(RingTheory/Etale): etale <=> separable for field extensions of finite type. (#17633) Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- Mathlib.lean | 1 + Mathlib/RingTheory/Etale/Basic.lean | 3 + Mathlib/RingTheory/Etale/Field.lean | 151 +++++++++++++++++++++++ Mathlib/RingTheory/Smooth/Basic.lean | 3 + Mathlib/RingTheory/Unramified/Basic.lean | 3 + 5 files changed, 161 insertions(+) create mode 100644 Mathlib/RingTheory/Etale/Field.lean diff --git a/Mathlib.lean b/Mathlib.lean index 1c4fb5f69a188..4aea7fc619759 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4004,6 +4004,7 @@ import Mathlib.RingTheory.DualNumber import Mathlib.RingTheory.EisensteinCriterion import Mathlib.RingTheory.EssentialFiniteness import Mathlib.RingTheory.Etale.Basic +import Mathlib.RingTheory.Etale.Field import Mathlib.RingTheory.EuclideanDomain import Mathlib.RingTheory.Filtration import Mathlib.RingTheory.FiniteLength diff --git a/Mathlib/RingTheory/Etale/Basic.lean b/Mathlib/RingTheory/Etale/Basic.lean index 70ab05f56b544..554788c7bc998 100644 --- a/Mathlib/RingTheory/Etale/Basic.lean +++ b/Mathlib/RingTheory/Etale/Basic.lean @@ -84,6 +84,9 @@ theorem of_equiv [FormallyEtale R A] (e : A ≃ₐ[R] B) : FormallyEtale R B := FormallyEtale.iff_unramified_and_smooth.mpr ⟨FormallyUnramified.of_equiv e, FormallySmooth.of_equiv e⟩ +theorem iff_of_equiv (e : A ≃ₐ[R] B) : FormallyEtale R A ↔ FormallyEtale R B := + ⟨fun _ ↦ of_equiv e, fun _ ↦ of_equiv e.symm⟩ + end OfEquiv section Comp diff --git a/Mathlib/RingTheory/Etale/Field.lean b/Mathlib/RingTheory/Etale/Field.lean new file mode 100644 index 0000000000000..8305070f2015e --- /dev/null +++ b/Mathlib/RingTheory/Etale/Field.lean @@ -0,0 +1,151 @@ +/- +Copyright (c) 2024 Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Andrew Yang +-/ +import Mathlib.RingTheory.Etale.Basic +import Mathlib.RingTheory.Unramified.Field + +/-! +# Etale algebras over fields + +## Main results + +Let `K` be a field, `A` be a `K`-algebra and `L` be a field extension of `K`. + +- `Algebra.FormallyEtale.of_isSeparable`: + If `L` is separable over `K`, then `L` is formally etale over `K`. +- `Algebra.FormallyEtale.iff_isSeparable`: + If `L` is (essentially) of finite type over `K`, then `L/K` is etale iff `L/K` is separable. + +## References + +- [B. Iversen, *Generic Local Structure of the Morphisms in Commutative Algebra*][iversen] + +-/ + + +universe u + +variable (K A L : Type u) [Field K] [Field L] [CommRing A] [Algebra K A] [Algebra K L] + +open Algebra Polynomial + +open scoped TensorProduct + +namespace Algebra.FormallyEtale + +/-- +This is a weaker version of `of_isSeparable` that additionally assumes `EssFiniteType K L`. +Use that instead. + +This is Iversen Corollary II.5.3. +-/ +theorem of_isSeparable_aux [Algebra.IsSeparable K L] [EssFiniteType K L] : + FormallyEtale K L := by + -- We already know that for field extensions + -- IsSeparable + EssFiniteType => FormallyUnramified + Finite + have := FormallyUnramified.of_isSeparable K L + have := FormallyUnramified.finite_of_free (R := K) (S := L) + constructor + -- We shall show that any `f : L → B/I` can be lifted to `L → B` if `I^2 = ⊥` + intros B _ _ I h + refine ⟨(FormallyUnramified.of_isSeparable K L).1 I h, ?_⟩ + intro f + -- By separability and finiteness, we may assume `L = K(α)` with `p` the minpoly of `α`. + let pb := Field.powerBasisOfFiniteOfSeparable K L + -- Let `x : B` such that `f(α) = x` in `B / I`. + obtain ⟨x, hx⟩ := Ideal.Quotient.mk_surjective (f pb.gen) + have helper : ∀ x, IsScalarTower.toAlgHom K B (B ⧸ I) x = Ideal.Quotient.mk I x := fun _ ↦ rfl + -- Then `p(x) = 0 mod I`, and the goal is to find some `ε ∈ I` such that + -- `p(x + ε) = p(x) + ε p'(x) = 0`, and we will get our lift into `B`. + have hx' : Ideal.Quotient.mk I (aeval x (minpoly K pb.gen)) = 0 := by + rw [← helper, ← aeval_algHom_apply, helper, hx, aeval_algHom_apply, minpoly.aeval, map_zero] + -- Since `p` is separable, `-p'(x)` is invertible in `B ⧸ I`, + obtain ⟨u, hu⟩ : ∃ u, (aeval x) (derivative (minpoly K pb.gen)) * u + 1 ∈ I := by + have := (isUnit_iff_ne_zero.mpr ((Algebra.IsSeparable.isSeparable K + pb.gen).aeval_derivative_ne_zero (minpoly.aeval K _))).map f + rw [← aeval_algHom_apply, ← hx, ← helper, aeval_algHom_apply, helper] at this + obtain ⟨u, hu⟩ := Ideal.Quotient.mk_surjective (-this.unit⁻¹ : B ⧸ I) + use u + rw [← Ideal.Quotient.eq_zero_iff_mem, map_add, map_mul, map_one, hu, mul_neg, + IsUnit.mul_val_inv, neg_add_cancel] + -- And `ε = p(x)/(-p'(x))` works. + use pb.liftEquiv.symm ⟨x + u * aeval x (minpoly K pb.gen), ?_⟩ + · apply pb.algHom_ext + simp [hx, hx'] + · rw [← eval_map_algebraMap, Polynomial.eval_add_of_sq_eq_zero, derivative_map, + ← one_mul (eval x _), eval_map_algebraMap, eval_map_algebraMap, ← mul_assoc, ← add_mul, + ← Ideal.mem_bot, ← h, pow_two, add_comm] + · exact Ideal.mul_mem_mul hu (Ideal.Quotient.eq_zero_iff_mem.mp hx') + rw [← Ideal.mem_bot, ← h] + apply Ideal.pow_mem_pow + rw [← Ideal.Quotient.eq_zero_iff_mem, map_mul, hx', mul_zero] + +open scoped IntermediateField in +lemma of_isSeparable [Algebra.IsSeparable K L] : FormallyEtale K L := by + constructor + intros B _ _ I h + -- We shall show that any `f : L → B/I` can be lifted to `L → B` if `I^2 = ⊥`. + -- But we already know that there exists a unique lift for every finite subfield of `L` + -- by `of_isSeparable_aux`, so we can glue them all together. + refine ⟨(FormallyUnramified.of_isSeparable K L).1 I h, ?_⟩ + intro f + have : ∀ k : L, ∃! g : K⟮k⟯ →ₐ[K] B, + (Ideal.Quotient.mkₐ K I).comp g = f.comp (IsScalarTower.toAlgHom K _ L) := by + intro k + have := IsSeparable.of_algHom _ _ (IsScalarTower.toAlgHom K (K⟮k⟯) L) + have := IntermediateField.adjoin.finiteDimensional + (Algebra.IsSeparable.isSeparable K k).isIntegral + have := FormallyEtale.of_isSeparable_aux K (K⟮k⟯) + have := FormallyEtale.comp_bijective (R := K) (A := K⟮k⟯) I h + exact this.existsUnique _ + choose g hg₁ hg₂ using this + have hg₃ : ∀ x y (h : x ∈ K⟮y⟯), g y ⟨x, h⟩ = g x (IntermediateField.AdjoinSimple.gen K x) := by + intro x y h + have e : K⟮x⟯ ≤ K⟮y⟯ := by + rw [IntermediateField.adjoin_le_iff] + rintro _ rfl + exact h + rw [← hg₂ _ ((g _).comp (IntermediateField.inclusion e))] + · rfl + apply AlgHom.ext + intro ⟨a, _⟩ + rw [← AlgHom.comp_assoc, hg₁, AlgHom.comp_assoc] + simp + have H : ∀ x y : L, ∃ α : L, x ∈ K⟮α⟯ ∧ y ∈ K⟮α⟯ := by + intro x y + have : FiniteDimensional K K⟮x, y⟯ := by + apply IntermediateField.finiteDimensional_adjoin + intro x _; exact (Algebra.IsSeparable.isSeparable K x).isIntegral + have := IsSeparable.of_algHom _ _ (IsScalarTower.toAlgHom K (K⟮x, y⟯) L) + obtain ⟨⟨α, hα⟩, e⟩ := Field.exists_primitive_element K K⟮x,y⟯ + apply_fun (IntermediateField.map (IntermediateField.val _)) at e + rw [IntermediateField.adjoin_map, ← AlgHom.fieldRange_eq_map] at e + simp only [IntermediateField.coe_val, Set.image_singleton, + IntermediateField.fieldRange_val] at e + have hx : x ∈ K⟮α⟯ := e ▸ IntermediateField.subset_adjoin K {x, y} (by simp) + have hy : y ∈ K⟮α⟯ := e ▸ IntermediateField.subset_adjoin K {x, y} (by simp) + exact ⟨α, hx, hy⟩ + refine ⟨⟨⟨⟨⟨fun x ↦ g x (IntermediateField.AdjoinSimple.gen K x), ?_⟩, ?_⟩, ?_, ?_⟩, ?_⟩, ?_⟩ + · show g 1 1 = 1; rw [map_one] + · intros x y + obtain ⟨α, hx, hy⟩ := H x y + simp only [← hg₃ _ _ hx, ← hg₃ _ _ hy, ← map_mul, ← hg₃ _ _ (mul_mem hx hy)] + rfl + · show g 0 0 = 0; rw [map_zero] + · intros x y + obtain ⟨α, hx, hy⟩ := H x y + simp only [← hg₃ _ _ hx, ← hg₃ _ _ hy, ← map_add, ← hg₃ _ _ (add_mem hx hy)] + rfl + · intro r + show g _ (algebraMap K _ r) = _ + rw [AlgHom.commutes] + · ext x + simpa using AlgHom.congr_fun (hg₁ x) (IntermediateField.AdjoinSimple.gen K x) + +theorem iff_isSeparable [EssFiniteType K L] : + FormallyEtale K L ↔ Algebra.IsSeparable K L := + ⟨fun _ ↦ FormallyUnramified.isSeparable K L, fun _ ↦ of_isSeparable K L⟩ + +end Algebra.FormallyEtale diff --git a/Mathlib/RingTheory/Smooth/Basic.lean b/Mathlib/RingTheory/Smooth/Basic.lean index 179e34269a3af..30f395c764148 100644 --- a/Mathlib/RingTheory/Smooth/Basic.lean +++ b/Mathlib/RingTheory/Smooth/Basic.lean @@ -142,6 +142,9 @@ theorem of_equiv [FormallySmooth R A] (e : A ≃ₐ[R] B) : FormallySmooth R B : rw [← AlgHom.comp_assoc, FormallySmooth.comp_lift, AlgHom.comp_assoc, AlgEquiv.comp_symm, AlgHom.comp_id] +theorem iff_of_equiv (e : A ≃ₐ[R] B) : FormallySmooth R A ↔ FormallySmooth R B := + ⟨fun _ ↦ of_equiv e, fun _ ↦ of_equiv e.symm⟩ + end OfEquiv section Polynomial diff --git a/Mathlib/RingTheory/Unramified/Basic.lean b/Mathlib/RingTheory/Unramified/Basic.lean index baf22e56e7160..cd3ce7546dad0 100644 --- a/Mathlib/RingTheory/Unramified/Basic.lean +++ b/Mathlib/RingTheory/Unramified/Basic.lean @@ -169,6 +169,9 @@ theorem of_equiv [FormallyUnramified R A] (e : A ≃ₐ[R] B) : FormallyUnramified R B := of_surjective e.toAlgHom e.surjective +theorem iff_of_equiv (e : A ≃ₐ[R] B) : FormallyUnramified R A ↔ FormallyUnramified R B := + ⟨fun _ ↦ of_equiv e, fun _ ↦ of_equiv e.symm⟩ + end of_surjective section BaseChange From 5e5ae835705f8a7daf5689cf86facc6a58baaa57 Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Sat, 26 Oct 2024 13:47:50 +0000 Subject: [PATCH 447/505] feat(NumberTheory/LSeries): changeLevel and Dirichlet L-functions (#17875) Relate the `L`-function of a character mod `M` and that of the corresponding imprimitive character modulo a multiple `N` of `M`. Co-authored-by: Michael Stoll --- .../SpecialFunctions/Pow/Continuity.lean | 4 + .../Analysis/SpecialFunctions/Pow/Deriv.lean | 10 +++ .../EulerProduct/DirichletLSeries.lean | 46 ++++++++++ .../LSeries/DirichletContinuation.lean | 87 ++++++++++++++++++- Mathlib/NumberTheory/LSeries/ZMod.lean | 10 +-- 5 files changed, 149 insertions(+), 8 deletions(-) diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Continuity.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Continuity.lean index 841d4fc4de757..e0046de2ae680 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Continuity.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Continuity.lean @@ -136,6 +136,10 @@ theorem ContinuousOn.cpow_const {b : ℂ} (hf : ContinuousOn f s) (h : ∀ a : α, a ∈ s → f a ∈ slitPlane) : ContinuousOn (fun x => f x ^ b) s := hf.cpow continuousOn_const h +@[fun_prop] +lemma continuous_const_cpow (z : ℂ) [NeZero z] : Continuous fun s : ℂ ↦ z ^ s := + continuous_id.const_cpow (.inl <| NeZero.ne z) + end CpowLimits section RpowLimits diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean index 02967ef06b371..7d78453430c3a 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean @@ -130,6 +130,16 @@ theorem Differentiable.const_cpow (hf : Differentiable ℂ f) (h0 : c ≠ 0 ∨ ∀ x, f x ≠ 0) : Differentiable ℂ (fun x ↦ c ^ f x) := fun x ↦ (hf x).const_cpow (h0.imp_right fun h ↦ h x) +@[fun_prop] +lemma differentiable_const_cpow_of_neZero (z : ℂ) [NeZero z] : + Differentiable ℂ fun s : ℂ ↦ z ^ s := + differentiable_id.const_cpow (.inl <| NeZero.ne z) + +@[fun_prop] +lemma differentiableAt_const_cpow_of_neZero (z : ℂ) [NeZero z] (t : ℂ) : + DifferentiableAt ℂ (fun s : ℂ ↦ z ^ s) t := + differentiableAt_id.const_cpow (.inl <| NeZero.ne z) + end fderiv section deriv diff --git a/Mathlib/NumberTheory/EulerProduct/DirichletLSeries.lean b/Mathlib/NumberTheory/EulerProduct/DirichletLSeries.lean index ba72ec28fdfcd..e3b05fad3ebd6 100644 --- a/Mathlib/NumberTheory/EulerProduct/DirichletLSeries.lean +++ b/Mathlib/NumberTheory/EulerProduct/DirichletLSeries.lean @@ -127,3 +127,49 @@ theorem dirichletLSeries_eulerProduct {N : ℕ} (χ : DirichletCharacter ℂ N) (𝓝 (L ↗χ s)) := by rw [← tsum_dirichletSummand χ hs] apply eulerProduct_completely_multiplicative <| summable_dirichletSummand χ hs + + +/-! +### Changing the level of a Dirichlet `L`-series +-/ + +/-- If `χ` is a Dirichlet character and its level `M` divides `N`, then we obtain the L-series +of `χ` considered as a Dirichlet character of level `N` from the L-series of `χ` by multiplying +with `∏ p ∈ N.primeFactors, (1 - χ p * p ^ (-s))`. -/ +lemma DirichletCharacter.LSeries_changeLevel {M N : ℕ} [NeZero N] + (hMN : M ∣ N) (χ : DirichletCharacter ℂ M) {s : ℂ} (hs : 1 < s.re) : + LSeries ↗(changeLevel hMN χ) s = + LSeries ↗χ s * ∏ p ∈ N.primeFactors, (1 - χ p * p ^ (-s)) := by + rw [prod_eq_tprod_mulIndicator, ← dirichletLSeries_eulerProduct_tprod _ hs, + ← dirichletLSeries_eulerProduct_tprod _ hs] + -- convert to a form suitable for `tprod_subtype` + have (f : Primes → ℂ) : ∏' (p : Primes), f p = ∏' (p : ↑{p : ℕ | p.Prime}), f p := rfl + rw [this, tprod_subtype _ fun p : ℕ ↦ (1 - (changeLevel hMN χ) p * p ^ (-s))⁻¹, + this, tprod_subtype _ fun p : ℕ ↦ (1 - χ p * p ^ (-s))⁻¹, ← tprod_mul] + rotate_left -- deal with convergence goals first + · exact multipliable_subtype_iff_mulIndicator.mp + (dirichletLSeries_eulerProduct_hasProd χ hs).multipliable + · exact multipliable_subtype_iff_mulIndicator.mp Multipliable.of_finite + · congr 1 with p + simp only [Set.mulIndicator_apply, Set.mem_setOf_eq, Finset.mem_coe, Nat.mem_primeFactors, + ne_eq, mul_ite, ite_mul, one_mul, mul_one] + by_cases h : p.Prime; swap + · simp only [h, false_and, if_false] + simp only [h, true_and, if_true] + by_cases hp' : p ∣ N; swap + · simp only [hp', false_and, ↓reduceIte, inv_inj, sub_right_inj, mul_eq_mul_right_iff, + cpow_eq_zero_iff, Nat.cast_eq_zero, h.ne_zero, ne_eq, neg_eq_zero, or_false] + have hq : IsUnit (p : ZMod N) := (ZMod.isUnit_prime_iff_not_dvd h).mpr hp' + simp only [hq.unit_spec ▸ DirichletCharacter.changeLevel_eq_cast_of_dvd χ hMN hq.unit, + ZMod.cast_natCast hMN] + · simp only [hp', NeZero.ne N, not_false_eq_true, and_self, ↓reduceIte] + have : ¬IsUnit (p : ZMod N) := by rwa [ZMod.isUnit_prime_iff_not_dvd h, not_not] + rw [MulChar.map_nonunit _ this, zero_mul, sub_zero, inv_one] + refine (inv_mul_cancel₀ ?_).symm + rw [sub_ne_zero, ne_comm] + -- Remains to show `χ p * p ^ (-s) ≠ 1`. We show its norm is strictly `< 1`. + apply_fun (‖·‖) + simp only [norm_mul, norm_one] + have ha : ‖χ p‖ ≤ 1 := χ.norm_le_one p + have hb : ‖(p : ℂ) ^ (-s)‖ ≤ 1 / 2 := norm_prime_cpow_le_one_half ⟨p, h⟩ hs + exact ((mul_le_mul ha hb (norm_nonneg _) zero_le_one).trans_lt (by norm_num)).ne diff --git a/Mathlib/NumberTheory/LSeries/DirichletContinuation.lean b/Mathlib/NumberTheory/LSeries/DirichletContinuation.lean index 2601a88444d1f..c6286c93a9c72 100644 --- a/Mathlib/NumberTheory/LSeries/DirichletContinuation.lean +++ b/Mathlib/NumberTheory/LSeries/DirichletContinuation.lean @@ -1,10 +1,11 @@ /- Copyright (c) 2024 David Loeffler. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: David Loeffler +Authors: David Loeffler, Michael Stoll -/ import Mathlib.NumberTheory.LSeries.ZMod import Mathlib.NumberTheory.DirichletCharacter.Basic +import Mathlib.NumberTheory.EulerProduct.DirichletLSeries /-! # Analytic continuation of Dirichlet L-functions @@ -37,9 +38,9 @@ All definitions and theorems are in the `DirichletCharacter` namespace. `completedLFunction χ s = N ^ (s - 1 / 2) * rootNumber χ * completedLFunction χ⁻¹ s`. -/ -open HurwitzZeta Complex Finset Classical ZMod +open HurwitzZeta Complex Finset Classical ZMod Filter -open scoped Real +open scoped Real Topology namespace DirichletCharacter @@ -103,6 +104,86 @@ lemma Even.LFunction_neg_two_mul_nat {χ : DirichletCharacter ℂ N} (hχ : Even LFunction χ (-(2 * n) - 1) = 0 := ZMod.LFunction_neg_two_mul_nat_sub_one hχ.to_fun n +/-! +## Results on changing levels +-/ + +private lemma LFunction_changeLevel_aux {M N : ℕ} [NeZero M] [NeZero N] (hMN : M ∣ N) + (χ : DirichletCharacter ℂ M) {s : ℂ} (hs : s ≠ 1) : + LFunction (changeLevel hMN χ) s = + LFunction χ s * ∏ p ∈ N.primeFactors, (1 - χ p * p ^ (-s)) := by + have hpc : IsPreconnected ({1}ᶜ : Set ℂ) := + (isConnected_compl_singleton_of_one_lt_rank (rank_real_complex ▸ Nat.one_lt_ofNat) _) + |>.isPreconnected + have hne : 2 ∈ ({1}ᶜ : Set ℂ) := by norm_num + refine AnalyticOnNhd.eqOn_of_preconnected_of_eventuallyEq (𝕜 := ℂ) + (g := fun s ↦ LFunction χ s * ∏ p ∈ N.primeFactors, (1 - χ p * p ^ (-s))) ?_ ?_ hpc hne ?_ hs + · refine DifferentiableOn.analyticOnNhd (fun s hs ↦ ?_) isOpen_compl_singleton + exact (differentiableAt_LFunction _ _ (.inl hs)).differentiableWithinAt + · refine DifferentiableOn.analyticOnNhd (fun s hs ↦ ?_) isOpen_compl_singleton + refine ((differentiableAt_LFunction _ _ (.inl hs)).mul ?_).differentiableWithinAt + refine .finset_prod fun i h ↦ ?_ + have : NeZero i := ⟨(Nat.pos_of_mem_primeFactors h).ne'⟩ + fun_prop + · refine eventually_of_mem ?_ (fun t (ht : 1 < t.re) ↦ ?_) + · exact (continuous_re.isOpen_preimage _ isOpen_Ioi).mem_nhds (by norm_num : 1 < (2 : ℂ).re) + · simpa only [LFunction_eq_LSeries _ ht] using LSeries_changeLevel hMN χ ht + +/-- If `χ` is a Dirichlet character and its level `M` divides `N`, then we obtain the L function +of `χ` considered as a Dirichlet character of level `N` from the L function of `χ` by multiplying +with `∏ p ∈ N.primeFactors, (1 - χ p * p ^ (-s))`. +(Note that `1 - χ p * p ^ (-s) = 1` when `p` divides `M`). -/ +lemma LFunction_changeLevel {M N : ℕ} [NeZero M] [NeZero N] (hMN : M ∣ N) + (χ : DirichletCharacter ℂ M) {s : ℂ} (h : χ ≠ 1 ∨ s ≠ 1) : + LFunction (changeLevel hMN χ) s = + LFunction χ s * ∏ p ∈ N.primeFactors, (1 - χ p * p ^ (-s)) := by + rcases h with h | h + · have hχ : changeLevel hMN χ ≠ 1 := h ∘ (changeLevel_eq_one_iff hMN).mp + have h' : Continuous fun s ↦ LFunction χ s * ∏ p ∈ N.primeFactors, (1 - χ p * ↑p ^ (-s)) := + (differentiable_LFunction h).continuous.mul <| continuous_finset_prod _ fun p hp ↦ by + have : NeZero p := ⟨(Nat.prime_of_mem_primeFactors hp).ne_zero⟩ + fun_prop + exact congrFun ((differentiable_LFunction hχ).continuous.ext_on + (dense_compl_singleton 1) h' (fun _ h ↦ LFunction_changeLevel_aux hMN χ h)) s + · exact LFunction_changeLevel_aux hMN χ h + +/-! +## The `L`-function of the trivial character mod `N` +-/ + +/-- The `L`-function of the trivial character mod `N`. -/ +noncomputable abbrev LFunctionTrivChar (N : ℕ) [NeZero N] := + (1 : DirichletCharacter ℂ N).LFunction + +/-- The L function of the trivial Dirichlet character mod `N` is obtained from the Riemann +zeta function by multiplying with `∏ p ∈ N.primeFactors, (1 - (p : ℂ) ^ (-s))`. -/ +lemma LFunctionTrivChar_eq_mul_riemannZeta {s : ℂ} (hs : s ≠ 1) : + LFunctionTrivChar N s = (∏ p ∈ N.primeFactors, (1 - (p : ℂ) ^ (-s))) * riemannZeta s := by + rw [← LFunction_modOne_eq (χ := 1), LFunctionTrivChar, ← changeLevel_one N.one_dvd, mul_comm] + convert LFunction_changeLevel N.one_dvd 1 (.inr hs) using 4 with p + rw [MulChar.one_apply <| isUnit_of_subsingleton _, one_mul] + +/-- The L function of the trivial Dirichlet character mod `N` has a simple pole with +residue `∏ p ∈ N.primeFactors, (1 - p⁻¹)` at `s = 1`. -/ +lemma LFunctionTrivChar_residue_one : + Tendsto (fun s ↦ (s - 1) * LFunctionTrivChar N s) (𝓝[≠] 1) + (𝓝 <| ∏ p ∈ N.primeFactors, (1 - (p : ℂ)⁻¹)) := by + have H : (fun s ↦ (s - 1) * LFunctionTrivChar N s) =ᶠ[𝓝[≠] 1] + fun s ↦ (∏ p ∈ N.primeFactors, (1 - (p : ℂ) ^ (-s))) * ((s - 1) * riemannZeta s) := by + refine Set.EqOn.eventuallyEq_nhdsWithin fun s hs ↦ ?_ + rw [mul_left_comm, LFunctionTrivChar_eq_mul_riemannZeta hs] + rw [tendsto_congr' H] + conv => enter [3, 1]; rw [← mul_one <| Finset.prod ..]; enter [1, 2, p]; rw [← cpow_neg_one] + refine .mul (f := fun s ↦ ∏ p ∈ N.primeFactors, _) ?_ riemannZeta_residue_one + refine tendsto_nhdsWithin_of_tendsto_nhds <| Continuous.tendsto ?_ 1 + exact continuous_finset_prod _ fun p hp ↦ by + have : NeZero p := ⟨(Nat.prime_of_mem_primeFactors hp).ne_zero⟩ + fun_prop + +/-! +## Completed L-functions and the functional equation +-/ + section gammaFactor omit [NeZero N] -- not required for these declarations diff --git a/Mathlib/NumberTheory/LSeries/ZMod.lean b/Mathlib/NumberTheory/LSeries/ZMod.lean index 9535bbf835aa8..a7b7dd924fbf6 100644 --- a/Mathlib/NumberTheory/LSeries/ZMod.lean +++ b/Mathlib/NumberTheory/LSeries/ZMod.lean @@ -118,7 +118,7 @@ lemma LFunction_eq_LSeries (Φ : ZMod N → ℂ) {s : ℂ} (hs : 1 < re s) : lemma differentiableAt_LFunction (Φ : ZMod N → ℂ) (s : ℂ) (hs : s ≠ 1 ∨ ∑ j, Φ j = 0) : DifferentiableAt ℂ (LFunction Φ) s := by - apply (differentiable_neg.const_cpow (Or.inl <| NeZero.ne _) s).mul + refine .mul (by fun_prop) ?_ rcases ne_or_eq s 1 with hs' | rfl · exact .sum fun j _ ↦ (differentiableAt_hurwitzZeta _ hs').const_mul _ · have := DifferentiableAt.sum (u := univ) fun j _ ↦ @@ -335,7 +335,7 @@ noncomputable def completedLFunction₀ (Φ : ZMod N → ℂ) (s : ℂ) : ℂ := lemma differentiable_completedLFunction₀ (Φ : ZMod N → ℂ) : Differentiable ℂ (completedLFunction₀ Φ) := by refine .add ?_ ?_ <;> - refine (differentiable_neg.const_cpow <| .inl <| NeZero.ne _).mul (.sum fun i _ ↦ .const_mul ?_ _) + refine .mul (by fun_prop) (.sum fun i _ ↦ .const_mul ?_ _) exacts [differentiable_completedHurwitzZetaEven₀ _, differentiable_completedHurwitzZetaOdd _] lemma completedLFunction_eq (Φ : ZMod N → ℂ) (s : ℂ) : @@ -357,12 +357,12 @@ lemma differentiableAt_completedLFunction (Φ : ZMod N → ℂ) (s : ℂ) (hs₀ -- correction terms from `completedLFunction_eq` are differentiable at `s`. refine ((differentiable_completedLFunction₀ _ _).sub ?_).sub ?_ · -- term with `1 / s` - refine ((differentiable_neg.const_cpow <| .inl <| NeZero.ne _) s).mul (hs₀.elim ?_ ?_) + refine .mul (by fun_prop) (hs₀.elim ?_ ?_) · exact fun h ↦ (differentiableAt_const _).div differentiableAt_id h · exact fun h ↦ by simp only [h, funext zero_div, differentiableAt_const] · -- term with `1 / (1 - s)` - refine ((differentiable_neg.const_cpow <| .inl <| NeZero.ne _) s).mul (hs₁.elim ?_ ?_) - · exact fun h ↦ (differentiableAt_const _).div (by fun_prop) (by rwa [sub_ne_zero, ne_comm]) + refine .mul (by fun_prop) (hs₁.elim ?_ ?_) + · exact fun h ↦ .div (by fun_prop) (by fun_prop) (by rwa [sub_ne_zero, ne_comm]) · exact fun h ↦ by simp only [h, zero_div, differentiableAt_const] /-- From d4e8628a6cac9e4d7e891cbb6bc9d6f258298f21 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Sat, 26 Oct 2024 13:47:52 +0000 Subject: [PATCH 448/505] feat(AlgebraicGeometry): morphisms into separated schemes are determined by a dense subset (#18151) Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- Mathlib.lean | 1 + .../Morphisms/Separated.lean | 55 +++++++++++++++++-- .../Morphisms/UnderlyingMap.lean | 13 +++++ Mathlib/CategoryTheory/IsConnected.lean | 4 +- .../Limits/Constructions/Over/Connected.lean | 4 +- .../Limits/Constructions/Over/Products.lean | 46 +++++++++++++++- .../Limits/Shapes/Connected.lean | 35 ++++++++++++ .../Limits/Shapes/FiniteLimits.lean | 5 +- 8 files changed, 149 insertions(+), 14 deletions(-) create mode 100644 Mathlib/CategoryTheory/Limits/Shapes/Connected.lean diff --git a/Mathlib.lean b/Mathlib.lean index 4aea7fc619759..802e7e7de1d6b 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1725,6 +1725,7 @@ import Mathlib.CategoryTheory.Limits.Shapes.BinaryProducts import Mathlib.CategoryTheory.Limits.Shapes.Biproducts import Mathlib.CategoryTheory.Limits.Shapes.CombinedProducts import Mathlib.CategoryTheory.Limits.Shapes.ConcreteCategory +import Mathlib.CategoryTheory.Limits.Shapes.Connected import Mathlib.CategoryTheory.Limits.Shapes.Countable import Mathlib.CategoryTheory.Limits.Shapes.Diagonal import Mathlib.CategoryTheory.Limits.Shapes.DisjointCoproduct diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Separated.lean b/Mathlib/AlgebraicGeometry/Morphisms/Separated.lean index 7f44de364381e..42ba861f7546f 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Separated.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Separated.lean @@ -4,9 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Christian Merten, Andrew Yang -/ import Mathlib.AlgebraicGeometry.Morphisms.ClosedImmersion -import Mathlib.AlgebraicGeometry.Morphisms.QuasiSeparated -import Mathlib.AlgebraicGeometry.Pullbacks -import Mathlib.CategoryTheory.MorphismProperty.Limits +import Mathlib.CategoryTheory.Limits.Constructions.Over.Basic +import Mathlib.CategoryTheory.Limits.Constructions.Over.Products import Mathlib.CategoryTheory.Limits.Shapes.Pullback.Equalizer /-! @@ -33,7 +32,7 @@ open scoped AlgebraicGeometry namespace AlgebraicGeometry -variable {X Y Z : Scheme.{u}} (f : X ⟶ Y) (g : Y ⟶ Z) +variable {W X Y Z : Scheme.{u}} (f : X ⟶ Y) (g : Y ⟶ Z) /-- A morphism is separated if the diagonal map is a closed immersion. -/ @[mk_iff] @@ -85,8 +84,8 @@ instance (R S : CommRingCat.{u}) (f : R ⟶ S) : IsSeparated (Spec.map f) := by exact .spec_of_surjective _ fun x ↦ ⟨.tmul R 1 x, (Algebra.TensorProduct.lmul'_apply_tmul (R := R) (S := S) 1 x).trans (one_mul x)⟩ -instance (priority := 100) [h : IsAffineHom f] : IsSeparated f := by - clear g Z +@[instance 100] +lemma of_isAffineHom [h : IsAffineHom f] : IsSeparated f := by wlog hY : IsAffine Y · rw [IsLocalAtTarget.iff_of_iSup_eq_top (P := @IsSeparated) _ (iSup_affineOpens_eq_top Y)] @@ -128,6 +127,40 @@ lemma IsSeparated.of_comp [IsSeparated (f ≫ g)] : IsSeparated f := by lemma IsSeparated.comp_iff [IsSeparated g] : IsSeparated (f ≫ g) ↔ IsSeparated f := ⟨fun _ ↦ .of_comp f g, fun _ ↦ inferInstance⟩ +@[stacks 01KM] +instance isClosedImmersion_equalizer_ι_left {S : Scheme} {X Y : Over S} [IsSeparated Y.hom] + (f g : X ⟶ Y) : IsClosedImmersion (equalizer.ι f g).left := by + refine IsClosedImmersion.stableUnderBaseChange + ((Limits.isPullback_equalizer_prod f g).map (Over.forget _)).flip ?_ + rw [← MorphismProperty.cancel_right_of_respectsIso @IsClosedImmersion _ + (Over.prodLeftIsoPullback Y Y).hom] + convert (inferInstanceAs (IsClosedImmersion (pullback.diagonal Y.hom))) + ext1 <;> simp [← Over.comp_left] + +/-- +Suppose `X` is a reduced scheme and that `f g : X ⟶ Y` agree over some separated `Y ⟶ Z`. +Then `f = g` if `ι ≫ f = ι ≫ g` for some dominant `ι`. +-/ +lemma ext_of_isDominant_of_isSeparated [IsReduced X] {f g : X ⟶ Y} + (s : Y ⟶ Z) [IsSeparated s] (h : f ≫ s = g ≫ s) + (ι : W ⟶ X) [IsDominant ι] (hU : ι ≫ f = ι ≫ g) : f = g := by + let X' : Over Z := Over.mk (f ≫ s) + let Y' : Over Z := Over.mk s + let U' : Over Z := Over.mk (ι ≫ f ≫ s) + let f' : X' ⟶ Y' := Over.homMk f + let g' : X' ⟶ Y' := Over.homMk g + let ι' : U' ⟶ X' := Over.homMk ι + have : IsSeparated Y'.hom := ‹_› + have : IsDominant (equalizer.ι f' g').left := by + apply (config := { allowSynthFailures := true }) IsDominant.of_comp (equalizer.lift ι' ?_).left + · rwa [← Over.comp_left, equalizer.lift_ι] + · ext1; exact hU + have : Surjective (equalizer.ι f' g').left := + surjective_of_isDominant_of_isClosed_range _ IsClosedImmersion.base_closed.2 + have := isIso_of_isClosedImmersion_of_surjective (Y := X) (equalizer.ι f' g').left + rw [← cancel_epi (equalizer.ι f' g').left] + exact congr($(equalizer.condition f' g').left) + namespace Scheme /-- A scheme `X` is separated if it is separated over `⊤_ Scheme`. -/ @@ -165,4 +198,14 @@ instance IsSeparated.hasAffineProperty : rw [Scheme.isSeparated_iff, ← terminal.comp_from f, IsSeparated.comp_iff] rfl +/-- +Suppose `f g : X ⟶ Y` where `X` is a reduced scheme and `Y` is a separated scheme. +Then `f = g` if `ι ≫ f = ι ≫ g` for some dominant `ι`. + +Also see `ext_of_isDominant_of_isSeparated` for the general version over arbitrary bases. +-/ +lemma ext_of_isDominant [IsReduced X] {f g : X ⟶ Y} [Y.IsSeparated] + (ι : W ⟶ X) [IsDominant ι] (hU : ι ≫ f = ι ≫ g) : f = g := + ext_of_isDominant_of_isSeparated (Limits.terminal.from _) (Limits.terminal.hom_ext _ _) ι hU + end AlgebraicGeometry diff --git a/Mathlib/AlgebraicGeometry/Morphisms/UnderlyingMap.lean b/Mathlib/AlgebraicGeometry/Morphisms/UnderlyingMap.lean index 3d31004a41fb1..aa7ef919ab2b8 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/UnderlyingMap.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/UnderlyingMap.lean @@ -176,6 +176,19 @@ instance IsDominant.isLocalAtTarget : IsLocalAtTarget @IsDominant := dominant_eq_topologically ▸ topologically_isLocalAtTarget' DenseRange fun _ _ _ hU _ ↦ denseRange_iff_denseRange_of_iSup_eq_top hU +lemma surjective_of_isDominant_of_isClosed_range (f : X ⟶ Y) [IsDominant f] + (hf : IsClosed (Set.range f.base)) : + Surjective f := + ⟨by rw [← Set.range_iff_surjective, ← hf.closure_eq, f.denseRange.closure_range]⟩ + +lemma IsDominant.of_comp_of_isOpenImmersion + (f : X ⟶ Y) (g : Y ⟶ Z) [H : IsDominant (f ≫ g)] [IsOpenImmersion g] : + IsDominant f := by + rw [isDominant_iff, DenseRange] at H ⊢ + simp only [Scheme.comp_coeBase, TopCat.coe_comp, Set.range_comp] at H + convert H.preimage g.isOpenEmbedding.isOpenMap using 1 + rw [Set.preimage_image_eq _ g.isOpenEmbedding.inj] + end IsDominant end AlgebraicGeometry diff --git a/Mathlib/CategoryTheory/IsConnected.lean b/Mathlib/CategoryTheory/IsConnected.lean index 6363451584e8a..c6b21cf4cbce6 100644 --- a/Mathlib/CategoryTheory/IsConnected.lean +++ b/Mathlib/CategoryTheory/IsConnected.lean @@ -412,7 +412,9 @@ def discreteIsConnectedEquivPUnit {α : Type u₁} [IsConnected (Discrete α)] : unitIso := isoConstant _ (Classical.arbitrary _) counitIso := Functor.punitExt _ _ } -variable {C : Type u₂} [Category.{u₁} C] +universe v u + +variable {C : Type u} [Category.{v} C] /-- For objects `X Y : C`, any natural transformation `α : const X ⟶ const Y` from a connected category must be constant. diff --git a/Mathlib/CategoryTheory/Limits/Constructions/Over/Connected.lean b/Mathlib/CategoryTheory/Limits/Constructions/Over/Connected.lean index 82bae7cbbc52d..18e56b5731467 100644 --- a/Mathlib/CategoryTheory/Limits/Constructions/Over/Connected.lean +++ b/Mathlib/CategoryTheory/Limits/Constructions/Over/Connected.lean @@ -15,14 +15,14 @@ any connected limit which `C` has. -/ -universe v u +universe v' u' v u -- morphism levels before object levels. See note [CategoryTheory universes]. noncomputable section open CategoryTheory CategoryTheory.Limits -variable {J : Type v} [SmallCategory J] +variable {J : Type u'} [Category.{v'} J] variable {C : Type u} [Category.{v} C] variable {X : C} diff --git a/Mathlib/CategoryTheory/Limits/Constructions/Over/Products.lean b/Mathlib/CategoryTheory/Limits/Constructions/Over/Products.lean index 873dd0fc1eda0..33d22f9978ac8 100644 --- a/Mathlib/CategoryTheory/Limits/Constructions/Over/Products.lean +++ b/Mathlib/CategoryTheory/Limits/Constructions/Over/Products.lean @@ -3,9 +3,7 @@ Copyright (c) 2018 Johan Commelin. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin, Reid Barton, Bhavik Mehta -/ -import Mathlib.CategoryTheory.Comma.Over -import Mathlib.CategoryTheory.Limits.Shapes.WidePullbacks -import Mathlib.CategoryTheory.Limits.Shapes.FiniteProducts +import Mathlib.CategoryTheory.Limits.Shapes.Pullback.CommSq /-! # Products in the over category @@ -166,4 +164,46 @@ theorem over_hasTerminal (B : C) : HasTerminal (Over B) where dsimp at this rwa [Category.comp_id, Category.comp_id] at this } } +section BinaryProduct + +variable {X : C} {Y Z : Over X} + +open Limits + +lemma isPullback_of_binaryFan_isLimit (c : BinaryFan Y Z) (hc : IsLimit c) : + IsPullback c.fst.left c.snd.left Y.hom Z.hom := + ⟨by simp, ⟨((IsLimit.postcomposeHomEquiv (diagramIsoCospan _) _).symm + ((IsLimit.ofConeEquiv (ConstructProducts.conesEquiv X _).symm).symm hc)).ofIsoLimit + (PullbackCone.isoMk _)⟩⟩ + +variable (Y Z) [HasPullback Y.hom Z.hom] [HasBinaryProduct Y Z] + +/-- The product of `Y` and `Z` in `Over X` is isomorpic to `Y ×ₓ Z`. -/ +noncomputable +def prodLeftIsoPullback : + (Y ⨯ Z).left ≅ pullback Y.hom Z.hom := + (Over.isPullback_of_binaryFan_isLimit _ (prodIsProd Y Z)).isoPullback + +@[reassoc (attr := simp)] +lemma prodLeftIsoPullback_hom_fst : + (prodLeftIsoPullback Y Z).hom ≫ pullback.fst _ _ = (prod.fst (X := Y)).left := + IsPullback.isoPullback_hom_fst _ + +@[reassoc (attr := simp)] +lemma prodLeftIsoPullback_hom_snd : + (prodLeftIsoPullback Y Z).hom ≫ pullback.snd _ _ = (prod.snd (X := Y)).left := + IsPullback.isoPullback_hom_snd _ + +@[reassoc (attr := simp)] +lemma prodLeftIsoPullback_inv_fst : + (prodLeftIsoPullback Y Z).inv ≫ (prod.fst (X := Y)).left = pullback.fst _ _ := + IsPullback.isoPullback_inv_fst _ + +@[reassoc (attr := simp)] +lemma prodLeftIsoPullback_inv_snd : + (prodLeftIsoPullback Y Z).inv ≫ (prod.snd (X := Y)).left = pullback.snd _ _ := + IsPullback.isoPullback_inv_snd _ + +end BinaryProduct + end CategoryTheory.Over diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Connected.lean b/Mathlib/CategoryTheory/Limits/Shapes/Connected.lean new file mode 100644 index 0000000000000..1755d71662292 --- /dev/null +++ b/Mathlib/CategoryTheory/Limits/Shapes/Connected.lean @@ -0,0 +1,35 @@ +/- +Copyright (c) 2024 Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Andrew Yang +-/ +import Mathlib.CategoryTheory.IsConnected +import Mathlib.CategoryTheory.Limits.Shapes.WidePullbacks + +/-! + +# Connected shapes + +In this file we prove that various shapes are connected. + +-/ + +namespace CategoryTheory + +open Limits + +instance {J} : IsConnected (WidePullbackShape J) := by + apply IsConnected.of_constant_of_preserves_morphisms + intros α F H + suffices ∀ i, F i = F none from fun j j' ↦ (this j).trans (this j').symm + rintro ⟨⟩ + exacts [rfl, H (.term _)] + +instance {J} : IsConnected (WidePushoutShape J) := by + apply IsConnected.of_constant_of_preserves_morphisms + intros α F H + suffices ∀ i, F i = F none from fun j j' ↦ (this j).trans (this j').symm + rintro ⟨⟩ + exacts [rfl, (H (.init _)).symm] + +end CategoryTheory diff --git a/Mathlib/CategoryTheory/Limits/Shapes/FiniteLimits.lean b/Mathlib/CategoryTheory/Limits/Shapes/FiniteLimits.lean index 812d28171a0d8..14a8a0a386b6a 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/FiniteLimits.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/FiniteLimits.lean @@ -239,13 +239,14 @@ instance hasColimitsOfShape_widePushoutShape (J : Type) [Finite J] [HasFiniteWid /-- Finite wide pullbacks are finite limits, so if `C` has all finite limits, it also has finite wide pullbacks -/ -theorem hasFiniteWidePullbacks_of_hasFiniteLimits [HasFiniteLimits C] : HasFiniteWidePullbacks C := +instance (priority := 900) hasFiniteWidePullbacks_of_hasFiniteLimits [HasFiniteLimits C] : + HasFiniteWidePullbacks C := ⟨fun J _ => by cases nonempty_fintype J; exact HasFiniteLimits.out _⟩ /-- Finite wide pushouts are finite colimits, so if `C` has all finite colimits, it also has finite wide pushouts -/ -theorem hasFiniteWidePushouts_of_has_finite_limits [HasFiniteColimits C] : +instance (priority := 900) hasFiniteWidePushouts_of_has_finite_limits [HasFiniteColimits C] : HasFiniteWidePushouts C := ⟨fun J _ => by cases nonempty_fintype J; exact HasFiniteColimits.out _⟩ From 60d9a7988ce099bcea1610e0cffffa6d4061b847 Mon Sep 17 00:00:00 2001 From: Riccardo Brasca Date: Sat, 26 Oct 2024 13:47:53 +0000 Subject: [PATCH 449/505] feat: add dvd_C_mul_X_sub_one_pow_add_one (#18225) From flt-regular --- Mathlib/Algebra/CharP/Lemmas.lean | 9 +++++++++ Mathlib/RingTheory/Polynomial/Cyclotomic/Basic.lean | 11 +++++++++++ 2 files changed, 20 insertions(+) diff --git a/Mathlib/Algebra/CharP/Lemmas.lean b/Mathlib/Algebra/CharP/Lemmas.lean index 271aa60a6d144..68a618f1adec5 100644 --- a/Mathlib/Algebra/CharP/Lemmas.lean +++ b/Mathlib/Algebra/CharP/Lemmas.lean @@ -238,6 +238,15 @@ lemma sub_pow_eq_mul_pow_sub_pow_div_char : sub_pow_eq_mul_pow_sub_pow_div_expChar .. end CharP + +lemma Nat.Prime.dvd_add_pow_sub_pow_of_dvd (hpri : p.Prime) {r : R} (h₁ : r ∣ x ^ p) + (h₂ : r ∣ p * x) : r ∣ (x + y) ^ p - y ^ p := by + rw [add_pow_prime_eq hpri, add_right_comm, add_assoc, add_sub_assoc, add_sub_cancel_right] + apply dvd_add h₁ (h₂.trans <| mul_dvd_mul_left _ <| Finset.dvd_sum _) + simp only [Finset.mem_Ioo, and_imp, mul_assoc] + intro i hi0 _ + exact dvd_mul_of_dvd_left (dvd_rfl.pow hi0.ne') _ + end CommRing diff --git a/Mathlib/RingTheory/Polynomial/Cyclotomic/Basic.lean b/Mathlib/RingTheory/Polynomial/Cyclotomic/Basic.lean index 1d3ab8d7d15d0..65b40dc6f9351 100644 --- a/Mathlib/RingTheory/Polynomial/Cyclotomic/Basic.lean +++ b/Mathlib/RingTheory/Polynomial/Cyclotomic/Basic.lean @@ -584,4 +584,15 @@ theorem orderOf_root_cyclotomic_dvd {n : ℕ} (hpos : 0 < n) {p : ℕ} [Fact p.P end Order +section miscellaneous + +lemma dvd_C_mul_X_sub_one_pow_add_one {R : Type*} [CommRing R] {p : ℕ} (hpri : p.Prime) + (hp : p ≠ 2) (a r : R) (h₁ : r ∣ a ^ p) (h₂ : r ∣ p * a) : C r ∣ (C a * X - 1) ^ p + 1 := by + have := hpri.dvd_add_pow_sub_pow_of_dvd (C a * X) (-1) (r := C r) ?_ ?_ + · rwa [← sub_eq_add_neg, (hpri.odd_of_ne_two hp).neg_pow, one_pow, sub_neg_eq_add] at this + · simp only [mul_pow, ← map_pow, dvd_mul_right, (_root_.map_dvd C h₁).trans] + simp only [map_mul, map_natCast, ← mul_assoc, dvd_mul_right, (_root_.map_dvd C h₂).trans] + +end miscellaneous + end Polynomial From 158ff637acbc9fbee3fe79af5ec451efe5749814 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Sat, 26 Oct 2024 13:47:54 +0000 Subject: [PATCH 450/505] chore(LinearAlgebra): rename unifEigenspace to genEigenspace (#18251) --- Mathlib/Algebra/Lie/Weights/Basic.lean | 8 +- Mathlib/Algebra/Lie/Weights/Linear.lean | 4 +- .../InnerProductSpace/JointEigenspace.lean | 2 +- Mathlib/LinearAlgebra/Eigenspace/Basic.lean | 504 ++++++++---------- Mathlib/LinearAlgebra/Eigenspace/Pi.lean | 10 +- .../LinearAlgebra/Eigenspace/Semisimple.lean | 28 +- .../Eigenspace/Triangularizable.lean | 90 ++-- Mathlib/LinearAlgebra/Eigenspace/Zero.lean | 2 +- 8 files changed, 287 insertions(+), 361 deletions(-) diff --git a/Mathlib/Algebra/Lie/Weights/Basic.lean b/Mathlib/Algebra/Lie/Weights/Basic.lean index a903c7e4c8c88..128294e766cf1 100644 --- a/Mathlib/Algebra/Lie/Weights/Basic.lean +++ b/Mathlib/Algebra/Lie/Weights/Basic.lean @@ -308,7 +308,7 @@ theorem exists_genWeightSpace_le_ker_of_isNoetherian [IsNoetherian R M] (χ : L intro m hm replace hm : m ∈ (toEnd R L M x).maxGenEigenspace (χ x) := genWeightSpace_le_genWeightSpaceOf M x χ hm - rwa [Module.End.maxGenEigenspace_eq, Module.End.genEigenspace_def] at hm + rwa [Module.End.maxGenEigenspace_eq, Module.End.genEigenspace_nat] at hm variable (R) in theorem exists_genWeightSpace_zero_le_ker_of_isNoetherian @@ -631,7 +631,7 @@ lemma disjoint_genWeightSpaceOf [NoZeroSMulDivisors R M] {x : L} {φ₁ φ₂ : Disjoint (genWeightSpaceOf M φ₁ x) (genWeightSpaceOf M φ₂ x) := by rw [LieSubmodule.disjoint_iff_coe_toSubmodule] dsimp [genWeightSpaceOf] - exact Module.End.disjoint_unifEigenspace _ h _ _ + exact Module.End.disjoint_genEigenspace _ h _ _ lemma disjoint_genWeightSpace [NoZeroSMulDivisors R M] {χ₁ χ₂ : L → R} (h : χ₁ ≠ χ₂) : Disjoint (genWeightSpace M χ₁) (genWeightSpace M χ₂) := by @@ -664,7 +664,7 @@ lemma independent_genWeightSpaceOf [NoZeroSMulDivisors R M] (x : L) : CompleteLattice.Independent fun (χ : R) ↦ genWeightSpaceOf M χ x := by rw [LieSubmodule.independent_iff_coe_toSubmodule] dsimp [genWeightSpaceOf] - exact (toEnd R L M x).independent_unifEigenspace _ + exact (toEnd R L M x).independent_genEigenspace _ lemma finite_genWeightSpaceOf_ne_bot [NoZeroSMulDivisors R M] [IsNoetherian R M] (x : L) : {χ : R | genWeightSpaceOf M χ x ≠ ⊥}.Finite := @@ -733,7 +733,7 @@ instance instIsTriangularizableOfIsAlgClosed [IsAlgClosed K] : IsTriangularizabl instance (N : LieSubmodule K L M) [IsTriangularizable K L M] : IsTriangularizable K L N := by refine ⟨fun y ↦ ?_⟩ rw [← N.toEnd_restrict_eq_toEnd y] - exact Module.End.unifEigenspace_restrict_eq_top _ (IsTriangularizable.maxGenEigenspace_eq_top y) + exact Module.End.genEigenspace_restrict_eq_top _ (IsTriangularizable.maxGenEigenspace_eq_top y) /-- For a triangularizable Lie module in finite dimensions, the weight spaces span the entire space. diff --git a/Mathlib/Algebra/Lie/Weights/Linear.lean b/Mathlib/Algebra/Lie/Weights/Linear.lean index 2898ac57b5d73..5e1934aa2fd39 100644 --- a/Mathlib/Algebra/Lie/Weights/Linear.lean +++ b/Mathlib/Algebra/Lie/Weights/Linear.lean @@ -99,13 +99,13 @@ instance instLinearWeightsOfIsLieAbelian [IsLieAbelian L] [NoZeroSMulDivisors R intro χ hχ x y simp_rw [Ne, ← LieSubmodule.coe_toSubmodule_eq_iff, genWeightSpace, genWeightSpaceOf, LieSubmodule.iInf_coe_toSubmodule, LieSubmodule.bot_coeSubmodule] at hχ - exact Module.End.map_add_of_iInf_unifEigenspace_ne_bot_of_commute + exact Module.End.map_add_of_iInf_genEigenspace_ne_bot_of_commute (toEnd R L M).toLinearMap χ _ hχ h x y { map_add := aux map_smul := fun χ hχ t x ↦ by simp_rw [Ne, ← LieSubmodule.coe_toSubmodule_eq_iff, genWeightSpace, genWeightSpaceOf, LieSubmodule.iInf_coe_toSubmodule, LieSubmodule.bot_coeSubmodule] at hχ - exact Module.End.map_smul_of_iInf_unifEigenspace_ne_bot + exact Module.End.map_smul_of_iInf_genEigenspace_ne_bot (toEnd R L M).toLinearMap χ _ hχ t x map_lie := fun χ hχ t x ↦ by rw [trivial_lie_zero, ← add_left_inj (χ 0), ← aux χ hχ, zero_add, zero_add] } diff --git a/Mathlib/Analysis/InnerProductSpace/JointEigenspace.lean b/Mathlib/Analysis/InnerProductSpace/JointEigenspace.lean index f10d670703e8a..049ba1532a238 100644 --- a/Mathlib/Analysis/InnerProductSpace/JointEigenspace.lean +++ b/Mathlib/Analysis/InnerProductSpace/JointEigenspace.lean @@ -72,7 +72,7 @@ theorem eigenspace_inf_eigenspace (hAB : A ∘ₗ B = B ∘ₗ A) (γ : 𝕜) : eigenspace A α ⊓ eigenspace B γ = map (Submodule.subtype (eigenspace A α)) (eigenspace (B.restrict (eigenspace_invariant_of_commute hAB α)) γ) := - (eigenspace A α).inf_unifEigenspace _ _ (k := 1) + (eigenspace A α).inf_genEigenspace _ _ (k := 1) variable [FiniteDimensional 𝕜 E] diff --git a/Mathlib/LinearAlgebra/Eigenspace/Basic.lean b/Mathlib/LinearAlgebra/Eigenspace/Basic.lean index 05a16ab184670..05c1324d366e9 100644 --- a/Mathlib/LinearAlgebra/Eigenspace/Basic.lean +++ b/Mathlib/LinearAlgebra/Eigenspace/Basic.lean @@ -58,29 +58,31 @@ 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 +/-- The submodule `genEigenspace 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 +(see Def 8.10 of [axler2015]), or the union of all these kernels if `k = ∞`. +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 := ⨆ 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 +lemma mem_genEigenspace {f : End R M} {μ : R} {k : ℕ∞} {x : M} : + x ∈ f.genEigenspace μ 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', + simp_rw [genEigenspace, 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 +lemma genEigenspace_directed {f : End R M} {μ : R} {k : ℕ∞} : + Directed (· ≤ ·) (fun l : {l : ℕ // l ≤ k} ↦ f.genEigenspace μ l) := by have aux : Monotone ((↑) : {l : ℕ // l ≤ k} → ℕ∞) := fun x y h ↦ by simpa using h - exact ((unifEigenspace f μ).monotone.comp aux).directed_le + exact ((genEigenspace 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] +lemma mem_genEigenspace_nat {f : End R M} {μ : R} {k : ℕ} {x : M} : + x ∈ f.genEigenspace μ k ↔ x ∈ LinearMap.ker ((f - μ • 1) ^ k) := by + rw [mem_genEigenspace] constructor · rintro ⟨l, hl, hx⟩ simp only [Nat.cast_le] at hl @@ -88,64 +90,64 @@ lemma mem_unifEigenspace_nat {f : End R M} {μ : R} {k : ℕ} {x : M} : · 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 mem_genEigenspace_top {f : End R M} {μ : R} {x : M} : + x ∈ f.genEigenspace μ ⊤ ↔ ∃ k : ℕ, x ∈ LinearMap.ker ((f - μ • 1) ^ k) := by + simp [mem_genEigenspace] -lemma unifEigenspace_nat {f : End R M} {μ : R} {k : ℕ} : - f.unifEigenspace μ k = LinearMap.ker ((f - μ • 1) ^ k) := by - ext; simp [mem_unifEigenspace_nat] +lemma genEigenspace_nat {f : End R M} {μ : R} {k : ℕ} : + f.genEigenspace μ k = LinearMap.ker ((f - μ • 1) ^ k) := by + ext; simp [mem_genEigenspace_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 genEigenspace_eq_iSup_genEigenspace_nat (f : End R M) (μ : R) (k : ℕ∞) : + f.genEigenspace μ k = ⨆ l : {l : ℕ // l ≤ k}, f.genEigenspace μ l := by + simp_rw [genEigenspace_nat, genEigenspace, 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] +lemma genEigenspace_top (f : End R M) (μ : R) : + f.genEigenspace μ ⊤ = ⨆ k : ℕ, f.genEigenspace μ k := by + rw [genEigenspace_eq_iSup_genEigenspace_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] +lemma genEigenspace_one {f : End R M} {μ : R} : + f.genEigenspace μ 1 = LinearMap.ker (f - μ • 1) := by + rw [← Nat.cast_one, genEigenspace_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, +lemma mem_genEigenspace_one {f : End R M} {μ : R} {x : M} : + x ∈ f.genEigenspace μ 1 ↔ f x = μ • x := by + rw [genEigenspace_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` can prove this using `genEigenspace_zero` +lemma mem_genEigenspace_zero {f : End R M} {μ : R} {x : M} : + x ∈ f.genEigenspace μ 0 ↔ x = 0 := by + rw [← Nat.cast_zero, mem_genEigenspace_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 +lemma genEigenspace_zero {f : End R M} {μ : R} : + f.genEigenspace μ 0 = ⊥ := by + ext; apply mem_genEigenspace_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] +lemma genEigenspace_zero_nat (f : End R M) (k : ℕ) : + f.genEigenspace 0 k = LinearMap.ker (f ^ k) := by + ext; simp [mem_genEigenspace_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`. +`x ∈ f.genEigenspace μ 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 + x ∈ f.genEigenspace μ 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 ≠ ⊥`. +`f.genEigenspace μ k ≠ ⊥`. For `k = 1`, this means that `μ` is an eigenvalue of `f`. -/ def HasUnifEigenvalue (f : End R M) (μ : R) (k : ℕ∞) : Prop := - f.unifEigenspace μ k ≠ ⊥ + f.genEigenspace μ 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 @@ -174,7 +176,7 @@ lemma HasUnifEigenvector.hasUnifEigenvalue {f : End R M} {μ : R} {k : ℕ∞} { 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 + mem_genEigenspace_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 @@ -189,7 +191,7 @@ lemma HasUnifEigenvalue.pow {f : End R M} {μ : R} (h : f.HasUnifEigenvalue μ 1 (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⟩ + exact ⟨m, by simpa [mem_genEigenspace_one] using hm.pow_apply n, hm.2⟩ /-- A nilpotent endomorphism has nilpotent eigenvalues. @@ -212,16 +214,16 @@ lemma HasUnifEigenvalue.mem_spectrum {f : End R M} {μ : R} (hμ : HasUnifEigenv 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] + HasUnifEigenvalue, genEigenspace_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) := +lemma genEigenspace_div (f : End K V) (a b : K) (hb : b ≠ 0) : + genEigenspace 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] + genEigenspace f (a / b) 1 = genEigenspace f (b⁻¹ * a) 1 := by rw [div_eq_mul_inv, mul_comm] + _ = LinearMap.ker (f - (b⁻¹ * a) • 1) := by rw [genEigenspace_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] @@ -229,13 +231,13 @@ lemma unifEigenspace_div (f : End K V) (a b : K) (hb : b ≠ 0) : /-- 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 := +def genEigenrange (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 +lemma genEigenrange_nat {f : End R M} {μ : R} {k : ℕ} : + f.genEigenrange μ k = LinearMap.range ((f - μ • 1) ^ k) := by ext x - simp only [unifEigenrange, Nat.cast_le, Submodule.mem_iInf, LinearMap.mem_range] + simp only [genEigenrange, Nat.cast_le, Submodule.mem_iInf, LinearMap.mem_range] constructor · intro h exact h _ le_rfl @@ -248,41 +250,41 @@ lemma unifEigenrange_nat {f : End R M} {μ : R} {k : ℕ} : 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 + simp [HasUnifEigenvalue, Nat.cast_zero, genEigenspace_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 + monotonicSequenceLimitIndex <| (f.genEigenspace μ).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 +lemma genEigenspace_top_eq_maxUnifEigenspaceIndex [h : IsNoetherian R M] (f : End R M) (μ : R) : + genEigenspace f μ ⊤ = f.genEigenspace μ (maxUnifEigenspaceIndex f μ) := by rw [isNoetherian_iff] at h have := WellFounded.iSup_eq_monotonicSequenceLimit h <| - (f.unifEigenspace μ).comp <| WithTop.coeOrderHom.toOrderHom + (f.genEigenspace μ).comp <| WithTop.coeOrderHom.toOrderHom convert this using 1 - simp only [unifEigenspace, OrderHom.coe_mk, le_top, iSup_pos, OrderHom.comp_coe, + simp only [genEigenspace, 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) +lemma genEigenspace_le_genEigenspace_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 + f.genEigenspace μ k ≤ f.genEigenspace μ (maxUnifEigenspaceIndex f μ) := by + rw [← genEigenspace_top_eq_maxUnifEigenspaceIndex] + exact (f.genEigenspace μ).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] +theorem genEigenspace_eq_genEigenspace_maxUnifEigenspaceIndex_of_le [IsNoetherian R M] (f : End R M) (μ : R) {k : ℕ} (hk : maxUnifEigenspaceIndex f μ ≤ k) : - f.unifEigenspace μ k = f.unifEigenspace μ (maxUnifEigenspaceIndex f μ) := + f.genEigenspace μ k = f.genEigenspace μ (maxUnifEigenspaceIndex f μ) := le_antisymm - (unifEigenspace_le_unifEigenspace_maxUnifEigenspaceIndex _ _ _) - ((f.unifEigenspace μ).monotone <| by simpa using hk) + (genEigenspace_le_genEigenspace_maxUnifEigenspaceIndex _ _ _) + ((f.genEigenspace μ).monotone <| by simpa using hk) /-- A generalized eigenvalue for some exponent `k` is also a generalized eigenvalue for exponents larger than `k`. -/ @@ -292,7 +294,7 @@ lemma HasUnifEigenvalue.le {f : End R M} {μ : R} {k m : ℕ∞} unfold HasUnifEigenvalue at * contrapose! hk rw [← le_bot_iff, ← hk] - exact (f.unifEigenspace _).monotone hm + exact (f.genEigenspace _).monotone hm /-- A generalized eigenvalue for some exponent `k` is also a generalized eigenvalue for positive exponents. -/ @@ -301,10 +303,10 @@ lemma HasUnifEigenvalue.lt {f : End R M} {μ : R} {k m : ℕ∞} 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 [genEigenspace_one, LinearMap.ker_eq_bot] at contra rw [eq_bot_iff] intro x hx - rw [mem_unifEigenspace] at hx + rw [mem_genEigenspace] at hx rcases hx with ⟨l, -, hx⟩ rwa [LinearMap.ker_eq_bot.mpr] at hx rw [LinearMap.coe_pow (f - μ • 1) l] @@ -321,33 +323,33 @@ lemma maxUnifEigenspaceIndex_le_finrank [FiniteDimensional K V] (f : End K V) ( 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] + · exact (f.genEigenspace μ).monotone <| WithTop.coeOrderHom.monotone hn + · show (f.genEigenspace μ) n ≤ (f.genEigenspace μ) (finrank K V) + rw [genEigenspace_nat, genEigenspace_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 μ +lemma genEigenspace_le_genEigenspace_finrank [FiniteDimensional K V] (f : End K V) + (μ : K) (k : ℕ∞) : f.genEigenspace μ k ≤ f.genEigenspace μ (finrank K V) := by + calc f.genEigenspace μ k + ≤ f.genEigenspace μ ⊤ := (f.genEigenspace _).monotone le_top + _ ≤ f.genEigenspace μ (finrank K V) := by + rw [genEigenspace_top_eq_maxUnifEigenspaceIndex] + exact (f.genEigenspace _).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] +theorem genEigenspace_eq_genEigenspace_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) := + f.genEigenspace μ k = f.genEigenspace μ (finrank K V) := le_antisymm - (unifEigenspace_le_unifEigenspace_finrank _ _ _) - ((f.unifEigenspace μ).monotone <| by simpa using hk) + (genEigenspace_le_genEigenspace_finrank _ _ _) + ((f.genEigenspace μ).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 +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 intro x hx - simp only [SetLike.mem_coe, mem_unifEigenspace, LinearMap.mem_ker] at hx ⊢ + simp only [SetLike.mem_coe, mem_genEigenspace, 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 @@ -356,41 +358,41 @@ lemma mapsTo_unifEigenspace_of_comm {f g : End R M} (h : Commute f g) (μ : R) ( 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 : ℕ) +lemma isNilpotent_restrict_genEigenspace_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) : + (f.genEigenspace μ k) (f.genEigenspace μ k) := + mapsTo_genEigenspace_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 [mem_genEigenspace_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) +lemma isNilpotent_restrict_genEigenspace_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 μ) μ _) : + (f.genEigenspace μ ⊤) (f.genEigenspace μ ⊤) := + mapsTo_genEigenspace_of_comm (Algebra.mul_sub_algebraMap_commutes f μ) μ _) : IsNilpotent ((f - μ • 1).restrict h) := by apply isNilpotent_restrict_of_le - on_goal 2 => apply isNilpotent_restrict_unifEigenspace_nat f μ (maxUnifEigenspaceIndex f μ) - rw [unifEigenspace_top_eq_maxUnifEigenspaceIndex] + on_goal 2 => apply isNilpotent_restrict_genEigenspace_nat f μ (maxUnifEigenspaceIndex f μ) + rw [genEigenspace_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])-/ abbrev eigenspace (f : End R M) (μ : R) : Submodule R M := - f.unifEigenspace μ 1 + f.genEigenspace μ 1 lemma eigenspace_def {f : End R M} {μ : R} : f.eigenspace μ = LinearMap.ker (f - μ • 1) := by - rw [eigenspace, unifEigenspace_one] + rw [eigenspace, genEigenspace_one] @[simp] 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] + simp only [eigenspace, ← Nat.cast_one (R := ℕ∞), genEigenspace_zero_nat, pow_one] /-- A nonzero element of an eigenspace is an eigenvector. (Def 5.7 of [axler2015]) -/ abbrev HasEigenvector (f : End R M) (μ : R) (x : M) : Prop := @@ -419,7 +421,7 @@ theorem hasEigenvalue_of_hasEigenvector {f : End R M} {μ : R} {x : M} (h : HasE h.hasUnifEigenvalue theorem mem_eigenspace_iff {f : End R M} {μ : R} {x : M} : x ∈ eigenspace f μ ↔ f x = μ • x := - mem_unifEigenspace_one + mem_genEigenspace_one nonrec theorem HasEigenvector.apply_eq_smul {f : End R M} {μ : R} {x : M} (hx : f.HasEigenvector μ x) : @@ -462,28 +464,12 @@ 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) := - 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 := f.unifEigenspace μ k - monotone' k l hkl := (f.unifEigenspace μ).monotone <| by simpa + genEigenspace_div f a b hb +@[deprecated genEigenspace_nat (since := "2024-10-28")] lemma genEigenspace_def (f : End R M) (μ : R) (k : ℕ) : - 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 := - mem_unifEigenspace_nat - -@[simp] -theorem genEigenspace_zero (f : End R M) (k : ℕ) : - f.genEigenspace 0 k = LinearMap.ker (f ^ k) := - unifEigenspace_zero_nat _ _ + f.genEigenspace μ k = LinearMap.ker ((f - μ • 1) ^ k) := + genEigenspace_nat /-- A nonzero element of a generalized eigenspace is a generalized eigenvector. (Def 8.9 of [axler2015])-/ @@ -501,14 +487,10 @@ abbrev HasGenEigenvalue (f : End R M) (μ : R) (k : ℕ) : Prop := 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`. -/ -abbrev genEigenrange (f : End R M) (μ : R) (k : ℕ) : Submodule R M := - unifEigenrange f μ k - +@[deprecated genEigenrange_nat (since := "2024-10-28")] lemma genEigenrange_def {f : End R M} {μ : R} {k : ℕ} : - f.genEigenrange μ k = LinearMap.range ((f - μ • 1) ^ k) := by - rw [genEigenrange, unifEigenrange_nat] + f.genEigenrange μ k = LinearMap.range ((f - μ • 1) ^ k) := + genEigenrange_nat /-- The exponent of a generalized eigenvalue is never 0. -/ theorem exp_ne_zero_of_hasGenEigenvalue {f : End R M} {μ : R} {k : ℕ} @@ -517,25 +499,25 @@ theorem exp_ne_zero_of_hasGenEigenvalue {f : End R M} {μ : R} {k : ℕ} /-- The union of the kernels of `(f - μ • id) ^ k` over all `k`. -/ abbrev maxGenEigenspace (f : End R M) (μ : R) : Submodule R M := - unifEigenspace f μ ⊤ + genEigenspace f μ ⊤ lemma iSup_genEigenspace_eq (f : End R M) (μ : R) : - ⨆ k, (f.genEigenspace μ) k = f.maxGenEigenspace μ := by - simp_rw [maxGenEigenspace, unifEigenspace_top, genEigenspace, OrderHom.coe_mk] + ⨆ k : ℕ, (f.genEigenspace μ) k = f.maxGenEigenspace μ := by + simp_rw [maxGenEigenspace, genEigenspace_top] @[deprecated iSup_genEigenspace_eq (since := "2024-10-23")] lemma maxGenEigenspace_def (f : End R M) (μ : R) : - f.maxGenEigenspace μ = ⨆ k, f.genEigenspace μ k := + f.maxGenEigenspace μ = ⨆ k : ℕ, f.genEigenspace μ k := (iSup_genEigenspace_eq f μ).symm theorem genEigenspace_le_maximal (f : End R M) (μ : R) (k : ℕ) : f.genEigenspace μ k ≤ f.maxGenEigenspace μ := - (f.unifEigenspace μ).monotone le_top + (f.genEigenspace μ).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 := - mem_unifEigenspace_top + mem_genEigenspace_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 @@ -547,7 +529,7 @@ noncomputable abbrev maxGenEigenspaceIndex (f : End R M) (μ : R) := `(f - μ • id) ^ k` for some `k`. -/ theorem maxGenEigenspace_eq [IsNoetherian R M] (f : End R M) (μ : R) : maxGenEigenspace f μ = f.genEigenspace μ (maxGenEigenspaceIndex f μ) := - unifEigenspace_top_eq_maxUnifEigenspaceIndex _ _ + genEigenspace_top_eq_maxUnifEigenspaceIndex _ _ /-- A generalized eigenvalue for some exponent `k` is also a generalized eigenvalue for exponents larger than `k`. -/ @@ -559,7 +541,7 @@ theorem hasGenEigenvalue_of_hasGenEigenvalue_of_le {f : End R M} {μ : R} {k : /-- 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.unifEigenspace _).monotone <| by simpa using Nat.succ_le_of_lt hk + (f.genEigenspace _).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) @@ -577,36 +559,20 @@ theorem hasGenEigenvalue_iff_hasEigenvalue {f : End R M} {μ : R} {k : ℕ} (hk f.HasGenEigenvalue μ k ↔ f.HasEigenvalue μ := 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) := - unifEigenspace_le_unifEigenspace_finrank _ _ _ - theorem maxGenEigenspace_eq_genEigenspace_finrank [FiniteDimensional K V] (f : End K V) (μ : K) : f.maxGenEigenspace μ = f.genEigenspace μ (finrank K V) := by - apply le_antisymm _ <| (f.unifEigenspace μ).monotone le_top - rw [unifEigenspace_top_eq_maxUnifEigenspaceIndex] + apply le_antisymm _ <| (f.genEigenspace μ).monotone le_top + rw [genEigenspace_top_eq_maxUnifEigenspaceIndex] apply genEigenspace_le_genEigenspace_finrank f μ -/-- Generalized eigenspaces for exponents at least `finrank K V` are equal to each other. -/ -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) := - 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) := - mapsTo_unifEigenspace_of_comm h μ k - lemma mapsTo_maxGenEigenspace_of_comm {f g : End R M} (h : Commute f g) (μ : R) : MapsTo g ↑(f.maxGenEigenspace μ) ↑(f.maxGenEigenspace μ) := - mapsTo_unifEigenspace_of_comm h μ ⊤ + mapsTo_genEigenspace_of_comm h μ ⊤ @[deprecated mapsTo_iSup_genEigenspace_of_comm (since := "2024-10-23")] 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 + MapsTo g ↑(⨆ k : ℕ, f.genEigenspace μ k) ↑(⨆ k : ℕ, f.genEigenspace μ k) := by rw [iSup_genEigenspace_eq] apply mapsTo_maxGenEigenspace_of_comm h @@ -616,7 +582,7 @@ lemma isNilpotent_restrict_sub_algebraMap (f : End R M) (μ : R) (k : ℕ) (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 _ _ _ + isNilpotent_restrict_genEigenspace_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) @@ -624,91 +590,85 @@ lemma isNilpotent_restrict_maxGenEigenspace_sub_algebraMap [IsNoetherian R M] (f ↑(f.maxGenEigenspace μ) ↑(f.maxGenEigenspace μ) := mapsTo_maxGenEigenspace_of_comm (Algebra.mul_sub_algebraMap_commutes f μ) μ) : IsNilpotent ((f - algebraMap R (End R M) μ).restrict h) := by - apply isNilpotent_restrict_of_le (q := f.unifEigenspace μ (maxUnifEigenspaceIndex f μ)) - _ (isNilpotent_restrict_unifEigenspace_nat f μ (maxUnifEigenspaceIndex f μ)) + apply isNilpotent_restrict_of_le (q := f.genEigenspace μ (maxUnifEigenspaceIndex f μ)) + _ (isNilpotent_restrict_genEigenspace_nat f μ (maxUnifEigenspaceIndex f μ)) rw [maxGenEigenspace_eq] - exact le_rfl set_option linter.deprecated false in /-- The restriction of `f - μ • 1` to the generalized `μ`-eigenspace is nilpotent. -/ @[deprecated isNilpotent_restrict_maxGenEigenspace_sub_algebraMap (since := "2024-10-23")] lemma isNilpotent_restrict_iSup_sub_algebraMap [IsNoetherian R M] (f : End R M) (μ : R) (h : MapsTo (f - algebraMap R (End R M) μ) - ↑(⨆ k, f.genEigenspace μ k) ↑(⨆ k, f.genEigenspace μ k) := + ↑(⨆ 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 - apply isNilpotent_restrict_of_le (q := f.unifEigenspace μ (maxUnifEigenspaceIndex f μ)) - _ (isNilpotent_restrict_unifEigenspace_nat f μ (maxUnifEigenspaceIndex f μ)) + apply isNilpotent_restrict_of_le (q := f.genEigenspace μ (maxUnifEigenspaceIndex f μ)) + _ (isNilpotent_restrict_genEigenspace_nat f μ (maxUnifEigenspaceIndex f μ)) apply iSup_le intro k - apply unifEigenspace_le_unifEigenspace_maxUnifEigenspaceIndex + apply genEigenspace_le_genEigenspace_maxUnifEigenspaceIndex -lemma disjoint_unifEigenspace [NoZeroSMulDivisors R M] +lemma disjoint_genEigenspace [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] + Disjoint (f.genEigenspace μ₁ k) (f.genEigenspace μ₂ l) := by + rw [genEigenspace_eq_iSup_genEigenspace_nat, genEigenspace_eq_iSup_genEigenspace_nat] + simp_rw [genEigenspace_directed.disjoint_iSup_left, genEigenspace_directed.disjoint_iSup_right] rintro ⟨k, -⟩ ⟨l, -⟩ nontriviality M have := NoZeroSMulDivisors.isReduced R M rw [disjoint_iff] - set p := f.unifEigenspace μ₁ k ⊓ f.unifEigenspace μ₂ l + set p := f.genEigenspace μ₁ k ⊓ f.genEigenspace μ₂ 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_unifEigenspace_of_comm (Algebra.mul_sub_algebraMap_commutes f μ₁) μ₁ k) - (mapsTo_unifEigenspace_of_comm (Algebra.mul_sub_algebraMap_commutes f μ₁) μ₂ l) + (mapsTo_genEigenspace_of_comm (Algebra.mul_sub_algebraMap_commutes f μ₁) μ₁ k) + (mapsTo_genEigenspace_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_unifEigenspace_of_comm (Algebra.mul_sub_algebraMap_commutes f μ₂) μ₁ k) - (mapsTo_unifEigenspace_of_comm (Algebra.mul_sub_algebraMap_commutes f μ₂) μ₂ l) + (mapsTo_genEigenspace_of_comm (Algebra.mul_sub_algebraMap_commutes f μ₂) μ₁ k) + (mapsTo_genEigenspace_of_comm (Algebra.mul_sub_algebraMap_commutes f μ₂) μ₂ l) have : IsNilpotent (f₂ - f₁) := by 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] - 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 + apply mapsTo_genEigenspace_of_comm (Algebra.mul_sub_algebraMap_commutes f _) + apply isNilpotent_restrict_genEigenspace_nat + apply mapsTo_genEigenspace_of_comm (Algebra.mul_sub_algebraMap_commutes f _) + apply isNilpotent_restrict_genEigenspace_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 injOn_unifEigenspace [NoZeroSMulDivisors R M] (f : End R M) (k : ℕ∞) : - InjOn (f.unifEigenspace · k) {μ | f.unifEigenspace μ k ≠ ⊥} := by +lemma injOn_genEigenspace [NoZeroSMulDivisors R M] (f : End R M) (k : ℕ∞) : + InjOn (f.genEigenspace · k) {μ | f.genEigenspace μ k ≠ ⊥} := by rintro μ₁ _ μ₂ hμ₂ hμ₁₂ by_contra contra apply hμ₂ - simpa only [hμ₁₂, disjoint_self] using f.disjoint_unifEigenspace contra k k + simpa only [hμ₁₂, disjoint_self] using f.disjoint_genEigenspace contra k k -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 - -@[deprecated disjoint_unifEigenspace (since := "2024-10-23")] +@[deprecated disjoint_genEigenspace (since := "2024-10-23")] lemma disjoint_iSup_genEigenspace [NoZeroSMulDivisors R M] (f : End R M) {μ₁ μ₂ : R} (hμ : μ₁ ≠ μ₂) : - Disjoint (⨆ k, f.genEigenspace μ₁ k) (⨆ k, f.genEigenspace μ₂ k) := by - simpa only [iSup_genEigenspace_eq] using disjoint_unifEigenspace f hμ ⊤ ⊤ + Disjoint (⨆ k : ℕ, f.genEigenspace μ₁ k) (⨆ k : ℕ, f.genEigenspace μ₂ k) := by + simpa only [iSup_genEigenspace_eq] using disjoint_genEigenspace f hμ ⊤ ⊤ lemma injOn_maxGenEigenspace [NoZeroSMulDivisors R M] (f : End R M) : InjOn (f.maxGenEigenspace ·) {μ | f.maxGenEigenspace μ ≠ ⊥} := - injOn_unifEigenspace f ⊤ + injOn_genEigenspace f ⊤ -@[deprecated injOn_unifEigenspace (since := "2024-10-23")] -lemma injOn_genEigenspace [NoZeroSMulDivisors R M] (f : End R M) : - InjOn (⨆ k, f.genEigenspace · k) {μ | ⨆ k, f.genEigenspace μ k ≠ ⊥} := by +@[deprecated injOn_genEigenspace (since := "2024-10-23")] +lemma injOn_iSup_genEigenspace [NoZeroSMulDivisors R M] (f : End R M) : + InjOn (⨆ k : ℕ, f.genEigenspace · k) {μ | ⨆ k : ℕ, f.genEigenspace μ k ≠ ⊥} := by simp_rw [iSup_genEigenspace_eq] apply injOn_maxGenEigenspace -theorem independent_unifEigenspace [NoZeroSMulDivisors R M] (f : End R M) (k : ℕ∞) : - CompleteLattice.Independent (f.unifEigenspace · k) := by +theorem independent_genEigenspace [NoZeroSMulDivisors R M] (f : End R M) (k : ℕ∞) : + CompleteLattice.Independent (f.genEigenspace · k) := by classical - suffices ∀ μ₁ (s : Finset R), μ₁ ∉ s → Disjoint (f.unifEigenspace μ₁ k) - (s.sup fun μ ↦ f.unifEigenspace μ k) by - simp_rw [CompleteLattice.independent_iff_supIndep_of_injOn (injOn_unifEigenspace f k), + suffices ∀ μ₁ (s : Finset R), μ₁ ∉ s → Disjoint (f.genEigenspace μ₁ k) + (s.sup fun μ ↦ f.genEigenspace μ k) by + simp_rw [CompleteLattice.independent_iff_supIndep_of_injOn (injOn_genEigenspace f k), Finset.supIndep_iff_disjoint_erase] exact fun s μ _ ↦ this _ _ (s.not_mem_erase μ) intro μ₁ s @@ -720,35 +680,35 @@ theorem independent_unifEigenspace [NoZeroSMulDivisors R M] (f : End R M) (k : rw [Finset.sup_insert, disjoint_iff, Submodule.eq_bot_iff] rintro x ⟨hx, hx'⟩ simp only [SetLike.mem_coe] at hx hx' - suffices x ∈ unifEigenspace f μ₂ k by - rw [← Submodule.mem_bot (R := R), ← (f.disjoint_unifEigenspace hμ₁₂ k k).eq_bot] + suffices x ∈ genEigenspace f μ₂ k by + rw [← Submodule.mem_bot (R := R), ← (f.disjoint_genEigenspace hμ₁₂ k k).eq_bot] exact ⟨hx, this⟩ obtain ⟨y, hy, z, hz, rfl⟩ := Submodule.mem_sup.mp hx'; clear hx' let g := f - μ₂ • 1 - simp_rw [mem_unifEigenspace, ← exists_prop] at hy ⊢ + simp_rw [mem_genEigenspace, ← exists_prop] at hy ⊢ peel hy with l hlk hl - simp only [mem_unifEigenspace_nat, LinearMap.mem_ker] at hl + simp only [mem_genEigenspace_nat, LinearMap.mem_ker] at hl have hyz : (g ^ l) (y + z) ∈ - (f.unifEigenspace μ₁ k) ⊓ s.sup fun μ ↦ f.unifEigenspace μ k := by - refine ⟨f.mapsTo_unifEigenspace_of_comm (g := g ^ l) ?_ μ₁ k hx, ?_⟩ + (f.genEigenspace μ₁ k) ⊓ s.sup fun μ ↦ f.genEigenspace μ k := by + refine ⟨f.mapsTo_genEigenspace_of_comm (g := g ^ l) ?_ μ₁ k hx, ?_⟩ · exact Algebra.mul_sub_algebraMap_pow_commutes f μ₂ l · rw [SetLike.mem_coe, map_add, hl, zero_add] - suffices (s.sup fun μ ↦ f.unifEigenspace μ k).map (g ^ l) ≤ - s.sup fun μ ↦ f.unifEigenspace μ k by exact this (Submodule.mem_map_of_mem hz) + suffices (s.sup fun μ ↦ f.genEigenspace μ k).map (g ^ l) ≤ + s.sup fun μ ↦ f.genEigenspace μ k by exact this (Submodule.mem_map_of_mem hz) simp_rw [Finset.sup_eq_iSup, Submodule.map_iSup (ι := R), Submodule.map_iSup (ι := _ ∈ s)] refine iSup₂_mono fun μ _ ↦ ?_ rintro - ⟨u, hu, rfl⟩ - refine f.mapsTo_unifEigenspace_of_comm ?_ μ k hu + refine f.mapsTo_genEigenspace_of_comm ?_ μ k hu exact Algebra.mul_sub_algebraMap_pow_commutes f μ₂ l rwa [ih.eq_bot, Submodule.mem_bot] at hyz theorem independent_maxGenEigenspace [NoZeroSMulDivisors R M] (f : End R M) : CompleteLattice.Independent f.maxGenEigenspace := by - apply independent_unifEigenspace + apply independent_genEigenspace -@[deprecated independent_unifEigenspace (since := "2024-10-23")] -theorem independent_genEigenspace [NoZeroSMulDivisors R M] (f : End R M) : - CompleteLattice.Independent (fun μ ↦ ⨆ k, f.genEigenspace μ k) := by +@[deprecated independent_genEigenspace (since := "2024-10-23")] +theorem independent_iSup_genEigenspace [NoZeroSMulDivisors R M] (f : End R M) : + CompleteLattice.Independent (fun μ ↦ ⨆ k : ℕ, f.genEigenspace μ k) := by simp_rw [iSup_genEigenspace_eq] apply independent_maxGenEigenspace @@ -756,7 +716,7 @@ theorem independent_genEigenspace [NoZeroSMulDivisors R M] (f : End R M) : any eigenspace has trivial intersection with the span of all the other eigenspaces. -/ theorem eigenspaces_independent [NoZeroSMulDivisors R M] (f : End R M) : CompleteLattice.Independent f.eigenspace := - (f.independent_unifEigenspace 1).mono fun _ ↦ le_rfl + (f.independent_genEigenspace 1).mono fun _ ↦ le_rfl /-- Eigenvectors corresponding to distinct eigenvalues of a linear operator are linearly independent. -/ @@ -779,39 +739,28 @@ theorem eigenvectors_linearIndependent [NoZeroSMulDivisors R M] /-- 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`. -/ -theorem genEigenspace_restrict (f : End R M) (p : Submodule R M) (k : ℕ) (μ : R) +theorem genEigenspace_restrict (f : End R M) (p : Submodule R M) (k : ℕ∞) (μ : R) (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_def, OrderHom.coe_mk, ← LinearMap.ker_comp] - induction' k with k ih + ext x + suffices ∀ l : ℕ, genEigenspace (LinearMap.restrict f hfp) μ l = + Submodule.comap p.subtype (f.genEigenspace μ l) by + simp_rw [mem_genEigenspace, ← mem_genEigenspace_nat, this, + Submodule.mem_comap, mem_genEigenspace (k := k), mem_genEigenspace_nat] + intro l + simp only [genEigenspace_nat, OrderHom.coe_mk, ← LinearMap.ker_comp] + induction' l with l ih · rw [pow_zero, pow_zero, LinearMap.one_eq_id] apply (Submodule.ker_subtype _).symm · erw [pow_succ, pow_succ, LinearMap.ker_comp, LinearMap.ker_comp, ih, ← LinearMap.ker_comp, LinearMap.comp_assoc] -/-- 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`. -/ -theorem unifEigenspace_restrict (f : End R M) (p : Submodule R M) (k : ℕ∞) (μ : R) - (hfp : ∀ x : M, x ∈ p → f x ∈ p) : - unifEigenspace (LinearMap.restrict f hfp) μ k = - Submodule.comap p.subtype (f.unifEigenspace μ k) := by - ext x - simp only [mem_unifEigenspace, LinearMap.mem_ker, ← mem_genEigenspace, genEigenspace_restrict, - Submodule.mem_comap] - -lemma _root_.Submodule.inf_unifEigenspace (f : End R M) (p : Submodule R M) {k : ℕ∞} {μ : R} - (hfp : ∀ x : M, x ∈ p → f x ∈ p) : - p ⊓ f.unifEigenspace μ k = - (unifEigenspace (LinearMap.restrict f hfp) μ k).map p.subtype := by - rw [f.unifEigenspace_restrict _ _ _ hfp, Submodule.map_comap_eq, Submodule.range_subtype] - -@[deprecated Submodule.inf_unifEigenspace (since := "2024-10-23")] -lemma _root_.Submodule.inf_genEigenspace (f : End R M) (p : Submodule R M) {k : ℕ} {μ : R} +lemma _root_.Submodule.inf_genEigenspace (f : End R M) (p : Submodule R M) {k : ℕ∞} {μ : R} (hfp : ∀ x : M, x ∈ p → f x ∈ p) : p ⊓ f.genEigenspace μ k = - (genEigenspace (LinearMap.restrict f hfp) μ k).map p.subtype := - Submodule.inf_unifEigenspace _ _ _ + (genEigenspace (LinearMap.restrict f hfp) μ k).map p.subtype := by + rw [f.genEigenspace_restrict _ _ _ hfp, Submodule.map_comap_eq, Submodule.range_subtype] 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} @@ -843,14 +792,13 @@ theorem generalized_eigenvec_disjoint_range_ker [FiniteDimensional K V] (f : End (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, 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 + rw [genEigenspace_nat, ← LinearMap.ker_comp]; rfl + _ = f.genEigenspace μ (finrank K V + finrank K V : ℕ) := by + simp_rw [← pow_add, genEigenspace_nat]; rfl _ = f.genEigenspace μ (finrank K V) := by 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] + rw [disjoint_iff_inf_le, genEigenrange_nat, LinearMap.range_eq_map, + Submodule.map_inf_eq_map_inf_comap, top_inf_eq, h, genEigenspace_nat] apply Submodule.map_comap_le /-- If an invariant subspace `p` of an endomorphism `f` is disjoint from the `μ`-eigenspace of `f`, @@ -869,7 +817,7 @@ theorem pos_finrank_genEigenspace_of_hasEigenvalue [FiniteDimensional K V] {f : 0 = finrank K (⊥ : Submodule K V) := by rw [finrank_bot] _ < finrank K (f.eigenspace μ) := Submodule.finrank_lt_finrank_of_lt (bot_lt_iff_ne_bot.2 hx) _ ≤ finrank K (f.genEigenspace μ k) := - Submodule.finrank_mono ((f.genEigenspace μ).monotone (Nat.succ_le_of_lt hk)) + Submodule.finrank_mono ((f.genEigenspace μ).monotone (by simpa using Nat.succ_le_of_lt hk)) /-- A linear map maps a generalized eigenrange into itself. -/ theorem map_genEigenrange_le {f : End K V} {μ : K} {n : ℕ} : @@ -877,33 +825,31 @@ 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, unifEigenrange_nat]; exact (LinearMap.range_comp _ _).symm + rw [genEigenrange_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 := by - rw [genEigenrange, unifEigenrange_nat] - apply LinearMap.map_le_range + _ ≤ f.genEigenrange μ n := by rw [genEigenrange_nat]; apply LinearMap.map_le_range -lemma unifEigenspace_le_smul (f : Module.End R M) (μ t : R) (k : ℕ∞) : - (f.unifEigenspace μ k) ≤ (t • f).unifEigenspace (t * μ) k := by +lemma genEigenspace_le_smul (f : Module.End R M) (μ t : R) (k : ℕ∞) : + (f.genEigenspace μ k) ≤ (t • f).genEigenspace (t * μ) k := by intro m hm - simp_rw [mem_unifEigenspace, ← exists_prop, LinearMap.mem_ker] at hm ⊢ + simp_rw [mem_genEigenspace, ← exists_prop, LinearMap.mem_ker] at hm ⊢ peel hm with l hlk hl rw [mul_smul, ← smul_sub, smul_pow, LinearMap.smul_apply, hl, smul_zero] -@[deprecated unifEigenspace_le_smul (since := "2024-10-23")] +@[deprecated genEigenspace_le_smul (since := "2024-10-23")] lemma iSup_genEigenspace_le_smul (f : Module.End R M) (μ t : R) : - (⨆ k, f.genEigenspace μ k) ≤ ⨆ k, (t • f).genEigenspace (t * μ) k := by + (⨆ k : ℕ, f.genEigenspace μ k) ≤ ⨆ k : ℕ, (t • f).genEigenspace (t * μ) k := by rw [iSup_genEigenspace_eq, iSup_genEigenspace_eq] - apply unifEigenspace_le_smul + apply genEigenspace_le_smul -lemma unifEigenspace_inf_le_add +lemma genEigenspace_inf_le_add (f₁ f₂ : End R M) (μ₁ μ₂ : R) (k₁ k₂ : ℕ∞) (h : Commute f₁ f₂) : - (f₁.unifEigenspace μ₁ k₁) ⊓ (f₂.unifEigenspace μ₂ k₂) ≤ - (f₁ + f₂).unifEigenspace (μ₁ + μ₂) (k₁ + k₂) := by + (f₁.genEigenspace μ₁ k₁) ⊓ (f₂.genEigenspace μ₂ k₂) ≤ + (f₁ + f₂).genEigenspace (μ₁ + μ₂) (k₁ + k₂) := by intro m hm - simp only [Submodule.mem_inf, mem_unifEigenspace, LinearMap.mem_ker] at hm ⊢ + simp only [Submodule.mem_inf, mem_genEigenspace, LinearMap.mem_ker] at hm ⊢ obtain ⟨⟨l₁, hlk₁, hl₁⟩, ⟨l₂, hlk₂, hl₂⟩⟩ := hm use l₁ + l₂ have : f₁ + f₂ - (μ₁ + μ₂) • 1 = (f₁ - μ₁ • 1) + (f₂ - μ₂ • 1) := by @@ -923,59 +869,59 @@ lemma unifEigenspace_inf_le_add LinearMap.map_zero] · rw [LinearMap.mul_apply, LinearMap.pow_map_zero_of_le hj hl₂, LinearMap.map_zero] -@[deprecated unifEigenspace_inf_le_add (since := "2024-10-23")] +@[deprecated genEigenspace_inf_le_add (since := "2024-10-23")] lemma iSup_genEigenspace_inf_le_add (f₁ f₂ : End R M) (μ₁ μ₂ : R) (h : Commute f₁ f₂) : - (⨆ k, f₁.genEigenspace μ₁ k) ⊓ (⨆ k, f₂.genEigenspace μ₂ k) ≤ - ⨆ k, (f₁ + f₂).genEigenspace (μ₁ + μ₂) k := by + (⨆ k : ℕ, f₁.genEigenspace μ₁ k) ⊓ (⨆ k : ℕ, f₂.genEigenspace μ₂ k) ≤ + ⨆ k : ℕ, (f₁ + f₂).genEigenspace (μ₁ + μ₂) k := by simp_rw [iSup_genEigenspace_eq] - apply unifEigenspace_inf_le_add + apply genEigenspace_inf_le_add assumption -lemma map_smul_of_iInf_unifEigenspace_ne_bot [NoZeroSMulDivisors R M] +lemma map_smul_of_iInf_genEigenspace_ne_bot [NoZeroSMulDivisors R M] {L F : Type*} [SMul R L] [FunLike F L (End R M)] [MulActionHomClass F R L (End R M)] (f : F) - (μ : L → R) (k : ℕ∞) (h_ne : ⨅ x, (f x).unifEigenspace (μ x) k ≠ ⊥) + (μ : L → R) (k : ℕ∞) (h_ne : ⨅ x, (f x).genEigenspace (μ x) k ≠ ⊥) (t : R) (x : L) : μ (t • x) = t • μ x := by by_contra contra - let g : L → Submodule R M := fun x ↦ (f x).unifEigenspace (μ x) k + let g : L → Submodule R M := fun x ↦ (f x).genEigenspace (μ x) k have : ⨅ x, g x ≤ g x ⊓ g (t • x) := le_inf_iff.mpr ⟨iInf_le g x, iInf_le g (t • x)⟩ refine h_ne <| eq_bot_iff.mpr (le_trans this (disjoint_iff_inf_le.mp ?_)) - apply Disjoint.mono_left (unifEigenspace_le_smul (f x) (μ x) t k) + apply Disjoint.mono_left (genEigenspace_le_smul (f x) (μ x) t k) simp only [g, map_smul] - exact disjoint_unifEigenspace (t • f x) (Ne.symm contra) k k + exact disjoint_genEigenspace (t • f x) (Ne.symm contra) k k -@[deprecated map_smul_of_iInf_unifEigenspace_ne_bot (since := "2024-10-23")] -lemma map_smul_of_iInf_genEigenspace_ne_bot [NoZeroSMulDivisors R M] +@[deprecated map_smul_of_iInf_genEigenspace_ne_bot (since := "2024-10-23")] +lemma map_smul_of_iInf_iSup_genEigenspace_ne_bot [NoZeroSMulDivisors R M] {L F : Type*} [SMul R L] [FunLike F L (End R M)] [MulActionHomClass F R L (End R M)] (f : F) - (μ : L → R) (h_ne : ⨅ x, ⨆ k, (f x).genEigenspace (μ x) k ≠ ⊥) + (μ : L → R) (h_ne : ⨅ x, ⨆ k : ℕ, (f x).genEigenspace (μ x) k ≠ ⊥) (t : R) (x : L) : μ (t • x) = t • μ x := by simp_rw [iSup_genEigenspace_eq] at h_ne - apply map_smul_of_iInf_unifEigenspace_ne_bot f μ ⊤ h_ne t x + apply map_smul_of_iInf_genEigenspace_ne_bot f μ ⊤ h_ne t x -lemma map_add_of_iInf_unifEigenspace_ne_bot_of_commute [NoZeroSMulDivisors R M] +lemma map_add_of_iInf_genEigenspace_ne_bot_of_commute [NoZeroSMulDivisors R M] {L F : Type*} [Add L] [FunLike F L (End R M)] [AddHomClass F L (End R M)] (f : F) - (μ : L → R) (k : ℕ∞) (h_ne : ⨅ x, (f x).unifEigenspace (μ x) k ≠ ⊥) + (μ : L → R) (k : ℕ∞) (h_ne : ⨅ x, (f x).genEigenspace (μ x) k ≠ ⊥) (h : ∀ x y, Commute (f x) (f y)) (x y : L) : μ (x + y) = μ x + μ y := by by_contra contra - let g : L → Submodule R M := fun x ↦ (f x).unifEigenspace (μ x) k + let g : L → Submodule R M := fun x ↦ (f x).genEigenspace (μ x) k have : ⨅ x, g x ≤ (g x ⊓ g y) ⊓ g (x + y) := le_inf_iff.mpr ⟨le_inf_iff.mpr ⟨iInf_le g x, iInf_le g y⟩, iInf_le g (x + y)⟩ refine h_ne <| eq_bot_iff.mpr (le_trans this (disjoint_iff_inf_le.mp ?_)) - apply Disjoint.mono_left (unifEigenspace_inf_le_add (f x) (f y) (μ x) (μ y) k k (h x y)) + apply Disjoint.mono_left (genEigenspace_inf_le_add (f x) (f y) (μ x) (μ y) k k (h x y)) simp only [g, map_add] - exact disjoint_unifEigenspace (f x + f y) (Ne.symm contra) _ k + exact disjoint_genEigenspace (f x + f y) (Ne.symm contra) _ k -@[deprecated map_add_of_iInf_unifEigenspace_ne_bot_of_commute (since := "2024-10-23")] -lemma map_add_of_iInf_genEigenspace_ne_bot_of_commute [NoZeroSMulDivisors R M] +@[deprecated map_add_of_iInf_genEigenspace_ne_bot_of_commute (since := "2024-10-23")] +lemma map_add_of_iInf_iSup_genEigenspace_ne_bot_of_commute [NoZeroSMulDivisors R M] {L F : Type*} [Add L] [FunLike F L (End R M)] [AddHomClass F L (End R M)] (f : F) - (μ : L → R) (h_ne : ⨅ x, ⨆ k, (f x).genEigenspace (μ x) k ≠ ⊥) + (μ : L → R) (h_ne : ⨅ x, ⨆ k : ℕ, (f x).genEigenspace (μ x) k ≠ ⊥) (h : ∀ x y, Commute (f x) (f y)) (x y : L) : μ (x + y) = μ x + μ y := by simp_rw [iSup_genEigenspace_eq] at h_ne - apply map_add_of_iInf_unifEigenspace_ne_bot_of_commute f μ ⊤ h_ne h x y + apply map_add_of_iInf_genEigenspace_ne_bot_of_commute f μ ⊤ h_ne h x y end End diff --git a/Mathlib/LinearAlgebra/Eigenspace/Pi.lean b/Mathlib/LinearAlgebra/Eigenspace/Pi.lean index 2980b51c48875..73c18265d61ab 100644 --- a/Mathlib/LinearAlgebra/Eigenspace/Pi.lean +++ b/Mathlib/LinearAlgebra/Eigenspace/Pi.lean @@ -45,7 +45,7 @@ lemma _root_.Submodule.inf_iInf_maxGenEigenspace_of_forall_mapsTo {μ : ι → R (⨅ i, maxGenEigenspace ((f i).restrict (hfp i)) (μ i)).map p.subtype := by cases isEmpty_or_nonempty ι · simp [iInf_of_isEmpty] - · simp_rw [inf_iInf, p.inf_unifEigenspace _ (hfp _), Submodule.map_iInf _ p.injective_subtype] + · simp_rw [inf_iInf, p.inf_genEigenspace _ (hfp _), Submodule.map_iInf _ p.injective_subtype] /-- Given a family of endomorphisms `i ↦ f i`, a family of candidate eigenvalues `i ↦ μ i`, and a distinguished index `i` whose maximal generalised `μ i`-eigenspace is invariant wrt every `f j`, @@ -66,14 +66,14 @@ lemma iInf_maxGenEigenspace_restrict_map_subtype_eq conv_rhs => enter [1] ext - rw [p.inf_unifEigenspace (f _) (h _)] + rw [p.inf_genEigenspace (f _) (h _)] variable [NoZeroSMulDivisors R M] lemma disjoint_iInf_maxGenEigenspace {χ₁ χ₂ : ι → R} (h : χ₁ ≠ χ₂) : Disjoint (⨅ i, (f i).maxGenEigenspace (χ₁ i)) (⨅ i, (f i).maxGenEigenspace (χ₂ i)) := by obtain ⟨j, hj⟩ : ∃ j, χ₁ j ≠ χ₂ j := Function.ne_iff.mp h - exact (End.disjoint_unifEigenspace (f j) hj ⊤ ⊤).mono (iInf_le _ j) (iInf_le _ j) + exact (End.disjoint_genEigenspace (f j) hj ⊤ ⊤).mono (iInf_le _ j) (iInf_le _ j) lemma injOn_iInf_maxGenEigenspace : InjOn (fun χ : ι → R ↦ ⨅ i, (f i).maxGenEigenspace (χ i)) @@ -167,7 +167,7 @@ lemma iSup_iInf_maxGenEigenspace_eq_top_of_forall_mapsTo [FiniteDimensional K M] 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.unifEigenspace_restrict_eq_top _ (h' j) + · exact fun j ↦ Module.End.genEigenspace_restrict_eq_top _ (h' j) · rfl replace ih (φ : K) : ⨆ (χ : ι → K) (_ : χ i = φ), ⨅ j, maxGenEigenspace ((f j).restrict (hi j φ)) (χ j) = ⊤ := by @@ -177,7 +177,7 @@ lemma iSup_iInf_maxGenEigenspace_eq_top_of_forall_mapsTo [FiniteDimensional K M] 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] - exact ((f i).disjoint_unifEigenspace hχ.symm _ _).mono_right (iInf_le _ i) + exact ((f i).disjoint_genEigenspace hχ.symm _ _).mono_right (iInf_le _ i) replace ih (φ : K) : ⨆ (χ : ι → K) (_ : χ i = φ), ⨅ j, maxGenEigenspace (f j) (χ j) = maxGenEigenspace (f i) φ := by diff --git a/Mathlib/LinearAlgebra/Eigenspace/Semisimple.lean b/Mathlib/LinearAlgebra/Eigenspace/Semisimple.lean index 26ccc18c96849..bf60fd45e813a 100644 --- a/Mathlib/LinearAlgebra/Eigenspace/Semisimple.lean +++ b/Mathlib/LinearAlgebra/Eigenspace/Semisimple.lean @@ -26,20 +26,20 @@ namespace Module.End variable {R M : Type*} [CommRing R] [AddCommGroup M] [Module R M] {f g : End R M} lemma apply_eq_of_mem_of_comm_of_isFinitelySemisimple_of_isNil - {μ : R} {k : ℕ∞} {m : M} (hm : m ∈ f.unifEigenspace μ k) + {μ : R} {k : ℕ∞} {m : M} (hm : m ∈ f.genEigenspace μ k) (hfg : Commute f g) (hss : g.IsFinitelySemisimple) (hnil : IsNilpotent (f - g)) : g m = μ • m := by - rw [f.mem_unifEigenspace] at hm + rw [f.mem_genEigenspace] at hm obtain ⟨l, -, hm⟩ := hm - rw [LinearMap.mem_ker, ← f.mem_genEigenspace] at hm + rw [← f.mem_genEigenspace_nat] at hm set p := f.genEigenspace μ l - have h₁ : MapsTo g p p := mapsTo_unifEigenspace_of_comm hfg μ l + have h₁ : MapsTo g p p := mapsTo_genEigenspace_of_comm hfg μ l have h₂ : MapsTo (g - algebraMap R (End R M) μ) p p := - mapsTo_unifEigenspace_of_comm (hfg.sub_right <| Algebra.commute_algebraMap_right μ f) μ l + mapsTo_genEigenspace_of_comm (hfg.sub_right <| Algebra.commute_algebraMap_right μ f) μ l have h₃ : MapsTo (f - g) p p := - mapsTo_unifEigenspace_of_comm (Commute.sub_right rfl hfg) μ l + mapsTo_genEigenspace_of_comm (Commute.sub_right rfl hfg) μ l have h₄ : MapsTo (f - algebraMap R (End R M) μ) p p := - mapsTo_unifEigenspace_of_comm (Algebra.mul_sub_algebraMap_commutes f μ) μ l + mapsTo_genEigenspace_of_comm (Algebra.mul_sub_algebraMap_commutes f μ) μ l replace hfg : Commute (f - algebraMap R (End R M) μ) (f - g) := (Commute.sub_right rfl hfg).sub_left <| Algebra.commute_algebraMap_left μ (f - g) suffices IsNilpotent ((g - algebraMap R (End R M) μ).restrict h₂) by @@ -49,23 +49,17 @@ lemma apply_eq_of_mem_of_comm_of_isFinitelySemisimple_of_isNil simpa [LinearMap.restrict_sub h₄ h₃] using (LinearMap.restrict_commute hfg h₄ h₃).isNilpotent_sub (f.isNilpotent_restrict_sub_algebraMap μ l) (Module.End.isNilpotent.restrict h₃ hnil) -lemma IsFinitelySemisimple.unifEigenspace_eq_eigenspace +lemma IsFinitelySemisimple.genEigenspace_eq_eigenspace (hf : f.IsFinitelySemisimple) (μ : R) {k : ℕ∞} (hk : 0 < k) : - f.unifEigenspace μ k = f.eigenspace μ := by - refine le_antisymm (fun m hm ↦ mem_eigenspace_iff.mpr ?_) (f.unifEigenspace μ |>.mono ?_) + f.genEigenspace μ k = f.eigenspace μ := by + refine le_antisymm (fun m hm ↦ mem_eigenspace_iff.mpr ?_) (f.genEigenspace μ |>.mono ?_) · apply apply_eq_of_mem_of_comm_of_isFinitelySemisimple_of_isNil hm rfl hf simp · exact Order.one_le_iff_pos.mpr hk -lemma IsFinitelySemisimple.genEigenspace_eq_eigenspace - (hf : f.IsFinitelySemisimple) (μ : R) {k : ℕ} (hk : 0 < k) : - f.genEigenspace μ k = f.eigenspace μ := by - refine le_antisymm (fun m hm ↦ mem_eigenspace_iff.mpr ?_) (eigenspace_le_genEigenspace hk) - exact apply_eq_of_mem_of_comm_of_isFinitelySemisimple_of_isNil hm rfl hf (by simp) - lemma IsFinitelySemisimple.maxGenEigenspace_eq_eigenspace (hf : f.IsFinitelySemisimple) (μ : R) : f.maxGenEigenspace μ = f.eigenspace μ := - hf.unifEigenspace_eq_eigenspace μ ENat.top_pos + hf.genEigenspace_eq_eigenspace μ ENat.top_pos end Module.End diff --git a/Mathlib/LinearAlgebra/Eigenspace/Triangularizable.lean b/Mathlib/LinearAlgebra/Eigenspace/Triangularizable.lean index 4ac57d9a9c04a..76112cc06a755 100644 --- a/Mathlib/LinearAlgebra/Eigenspace/Triangularizable.lean +++ b/Mathlib/LinearAlgebra/Eigenspace/Triangularizable.lean @@ -45,20 +45,20 @@ variable {K V : Type*} [Field K] [AddCommGroup V] [Module K V] namespace Module.End -theorem exists_hasEigenvalue_of_unifEigenspace_eq_top [Nontrivial M] {f : End R M} (k : ℕ∞) - (hf : ⨆ μ, f.unifEigenspace μ k = ⊤) : +theorem exists_hasEigenvalue_of_genEigenspace_eq_top [Nontrivial M] {f : End R M} (k : ℕ∞) + (hf : ⨆ μ, f.genEigenspace μ k = ⊤) : ∃ μ, f.HasEigenvalue μ := by suffices ∃ μ, f.HasUnifEigenvalue μ k by peel this with μ hμ exact HasUnifEigenvalue.lt zero_lt_one hμ simp [HasUnifEigenvalue, ← not_forall, ← iSup_eq_bot, hf] -@[deprecated exists_hasEigenvalue_of_unifEigenspace_eq_top (since := "2024-10-11")] +@[deprecated exists_hasEigenvalue_of_genEigenspace_eq_top (since := "2024-10-11")] theorem exists_hasEigenvalue_of_iSup_genEigenspace_eq_top [Nontrivial M] {f : End R M} - (hf : ⨆ μ, ⨆ k, f.genEigenspace μ k = ⊤) : + (hf : ⨆ μ, ⨆ k : ℕ, f.genEigenspace μ k = ⊤) : ∃ μ, f.HasEigenvalue μ := by simp_rw [iSup_genEigenspace_eq] at hf - apply exists_hasEigenvalue_of_unifEigenspace_eq_top _ hf + apply exists_hasEigenvalue_of_genEigenspace_eq_top _ hf -- This is Lemma 5.21 of [axler2015], although we are no longer following that proof. /-- In finite dimensions, over an algebraically closed field, every linear endomorphism has an @@ -104,7 +104,7 @@ theorem iSup_maxGenEigenspace_eq_top [IsAlgClosed K] [FiniteDimensional K V] (f -- and the dimensions of `ES` and `ER` add up to `finrank K V`. have h_dim_add : finrank K ER + finrank K ES = finrank K V := by dsimp only [ER, ES] - rw [Module.End.genEigenspace_def, Module.End.genEigenrange_def] + rw [Module.End.genEigenspace_nat, Module.End.genEigenrange_nat] apply LinearMap.finrank_range_add_finrank_ker -- Therefore the dimension `ER` mus be smaller than `finrank K V`. have h_dim_ER : finrank K ER < n.succ := by linarith @@ -120,7 +120,7 @@ theorem iSup_maxGenEigenspace_eq_top [IsAlgClosed K] [FiniteDimensional K V] (f have hff' : ∀ μ, (f'.maxGenEigenspace μ).map ER.subtype ≤ f.maxGenEigenspace μ := by intros - rw [maxGenEigenspace, unifEigenspace_restrict] + rw [maxGenEigenspace, genEigenspace_restrict] apply Submodule.map_comap_le -- It follows that `ER` is contained in the span of all generalized eigenvectors. have hER : ER ≤ ⨆ (μ : K), f.maxGenEigenspace μ := by @@ -128,7 +128,7 @@ theorem iSup_maxGenEigenspace_eq_top [IsAlgClosed K] [FiniteDimensional K V] (f exact iSup_mono hff' -- `ES` is contained in this span by definition. have hES : ES ≤ ⨆ (μ : K), f.maxGenEigenspace μ := - ((f.unifEigenspace μ₀).mono le_top).trans (le_iSup f.maxGenEigenspace μ₀) + ((f.genEigenspace μ₀).mono le_top).trans (le_iSup f.maxGenEigenspace μ₀) -- Moreover, we know that `ER` and `ES` are disjoint. have h_disjoint : Disjoint ER ES := generalized_eigenvec_disjoint_range_ker f μ₀ -- Since the dimensions of `ER` and `ES` add up to the dimension of `V`, it follows that the @@ -152,12 +152,12 @@ namespace Submodule variable {p : Submodule K V} {f : Module.End K V} -theorem inf_iSup_unifEigenspace [FiniteDimensional K V] (h : ∀ x ∈ p, f x ∈ p) (k : ℕ∞) : - p ⊓ ⨆ μ, f.unifEigenspace μ k = ⨆ μ, p ⊓ f.unifEigenspace μ k := by +theorem inf_iSup_genEigenspace [FiniteDimensional K V] (h : ∀ x ∈ p, f x ∈ p) (k : ℕ∞) : + p ⊓ ⨆ μ, f.genEigenspace μ k = ⨆ μ, p ⊓ f.genEigenspace μ k := by refine le_antisymm (fun m hm ↦ ?_) (le_inf_iff.mpr ⟨iSup_le fun μ ↦ inf_le_left, iSup_mono fun μ ↦ inf_le_right⟩) classical - obtain ⟨hm₀ : m ∈ p, hm₁ : m ∈ ⨆ μ, f.unifEigenspace μ k⟩ := hm + obtain ⟨hm₀ : m ∈ p, hm₁ : m ∈ ⨆ μ, f.genEigenspace μ k⟩ := hm obtain ⟨m, hm₂, rfl⟩ := (mem_iSup_iff_exists_finsupp _ _).mp hm₁ suffices ∀ μ, (m μ : V) ∈ p by exact (mem_iSup_iff_exists_finsupp _ _).mpr ⟨m, fun μ ↦ mem_inf.mp ⟨this μ, hm₂ μ⟩, rfl⟩ @@ -165,7 +165,7 @@ theorem inf_iSup_unifEigenspace [FiniteDimensional K V] (h : ∀ x ∈ p, f x by_cases hμ : μ ∈ m.support; swap · simp only [Finsupp.not_mem_support_iff.mp hμ, p.zero_mem] have hm₂_aux := hm₂ - simp_rw [Module.End.mem_unifEigenspace] at hm₂_aux + simp_rw [Module.End.mem_genEigenspace] at hm₂_aux choose l hlk hl using hm₂_aux let l₀ : ℕ := m.support.sup l have h_comm : ∀ (μ₁ μ₂ : K), @@ -184,9 +184,9 @@ theorem inf_iSup_unifEigenspace [FiniteDimensional K V] (h : ∀ x ∈ p, f x split_ifs with hμμ' · rw [hμμ'] have hl₀ : ((f - algebraMap K (End K V) μ') ^ l₀) (m μ') = 0 := by - rw [← LinearMap.mem_ker, Algebra.algebraMap_eq_smul_one, ← End.mem_unifEigenspace_nat] - simp_rw [← End.mem_unifEigenspace_nat] at hl - suffices (l μ' : ℕ∞) ≤ l₀ from (f.unifEigenspace μ').mono this (hl μ') + rw [← LinearMap.mem_ker, Algebra.algebraMap_eq_smul_one, ← End.mem_genEigenspace_nat] + simp_rw [← End.mem_genEigenspace_nat] at hl + suffices (l μ' : ℕ∞) ≤ l₀ from (f.genEigenspace μ').mono this (hl μ') simpa only [Nat.cast_le] using Finset.le_sup hμ' have : _ = g := (m.support.erase μ).noncommProd_erase_mul (Finset.mem_erase.mpr ⟨hμμ', hμ'⟩) (fun μ ↦ (f - algebraMap K (End K V) μ) ^ l₀) (fun μ₁ _ μ₂ _ _ ↦ h_comm μ₁ μ₂) @@ -198,14 +198,14 @@ theorem inf_iSup_unifEigenspace [FiniteDimensional K V] (h : ∀ x ∈ p, f x intro x hx rw [LinearMap.sub_apply, algebraMap_end_apply] exact p.sub_mem (h _ hx) (smul_mem p μ' hx) - have hg₂ : MapsTo g ↑(f.unifEigenspace μ k) ↑(f.unifEigenspace μ k) := - f.mapsTo_unifEigenspace_of_comm hfg μ k - have hg₃ : InjOn g ↑(f.unifEigenspace μ k) := by + have hg₂ : MapsTo g ↑(f.genEigenspace μ k) ↑(f.genEigenspace μ k) := + f.mapsTo_genEigenspace_of_comm hfg μ k + have hg₃ : InjOn g ↑(f.genEigenspace μ k) := by apply LinearMap.injOn_of_disjoint_ker (subset_refl _) - have this := f.independent_unifEigenspace k + have this := f.independent_genEigenspace k have aux (μ') (_hμ' : μ' ∈ m.support.erase μ) : - (f.unifEigenspace μ') ↑l₀ ≤ (f.unifEigenspace μ') k := by - apply (f.unifEigenspace μ').mono + (f.genEigenspace μ') ↑l₀ ≤ (f.genEigenspace μ') k := by + apply (f.genEigenspace μ').mono rintro k rfl simp only [ENat.some_eq_coe, Nat.cast_inj, exists_eq_left'] apply Finset.sup_le @@ -216,64 +216,50 @@ theorem inf_iSup_unifEigenspace [FiniteDimensional K V] (h : ∀ x ∈ p, f x apply this.mono_right apply Finset.sup_mono_fun intro μ' hμ' - rw [Algebra.algebraMap_eq_smul_one, ← End.unifEigenspace_nat] + rw [Algebra.algebraMap_eq_smul_one, ← End.genEigenspace_nat] apply aux μ' hμ' · have := this.supIndep' (m.support.erase μ) apply Finset.supIndep_antimono_fun _ this intro μ' hμ' - rw [Algebra.algebraMap_eq_smul_one, ← End.unifEigenspace_nat] + rw [Algebra.algebraMap_eq_smul_one, ← End.genEigenspace_nat] apply aux μ' hμ' have hg₄ : SurjOn g - ↑(p ⊓ f.unifEigenspace μ k) ↑(p ⊓ f.unifEigenspace μ k) := by + ↑(p ⊓ f.genEigenspace μ k) ↑(p ⊓ f.genEigenspace μ k) := by have : MapsTo g - ↑(p ⊓ f.unifEigenspace μ k) ↑(p ⊓ f.unifEigenspace μ k) := + ↑(p ⊓ f.genEigenspace μ k) ↑(p ⊓ f.genEigenspace μ k) := hg₁.inter_inter hg₂ rw [← LinearMap.injOn_iff_surjOn this] exact hg₃.mono inter_subset_right specialize hm₂ μ - obtain ⟨y, ⟨hy₀ : y ∈ p, hy₁ : y ∈ f.unifEigenspace μ k⟩, hy₂ : g y = g (m μ)⟩ := + obtain ⟨y, ⟨hy₀ : y ∈ p, hy₁ : y ∈ f.genEigenspace μ k⟩, hy₂ : g y = g (m μ)⟩ := hg₄ ⟨(hg₀ ▸ hg₁ hm₀), hg₂ hm₂⟩ rwa [← hg₃ hy₁ hm₂ hy₂] -set_option linter.deprecated false in -@[deprecated inf_iSup_unifEigenspace (since := "2024-10-11")] -theorem inf_iSup_genEigenspace [FiniteDimensional K V] (h : ∀ x ∈ p, f x ∈ p) : - p ⊓ ⨆ μ, ⨆ k, f.genEigenspace μ k = ⨆ μ, ⨆ k, p ⊓ f.genEigenspace μ k := by - simp_rw [← (f.genEigenspace _).mono.directed_le.inf_iSup_eq, f.iSup_genEigenspace_eq] - apply inf_iSup_unifEigenspace h ⊤ - -theorem eq_iSup_inf_unifEigenspace [FiniteDimensional K V] (k : ℕ∞) - (h : ∀ x ∈ p, f x ∈ p) (h' : ⨆ μ, f.unifEigenspace μ k = ⊤) : - p = ⨆ μ, p ⊓ f.unifEigenspace μ k := by - rw [← inf_iSup_unifEigenspace h, h', inf_top_eq] - -set_option linter.deprecated false in -@[deprecated eq_iSup_inf_unifEigenspace (since := "2024-10-11")] -theorem eq_iSup_inf_genEigenspace [FiniteDimensional K V] - (h : ∀ x ∈ p, f x ∈ p) (h' : ⨆ μ, ⨆ k, f.genEigenspace μ k = ⊤) : - p = ⨆ μ, ⨆ k, p ⊓ f.genEigenspace μ k := by +theorem eq_iSup_inf_genEigenspace [FiniteDimensional K V] (k : ℕ∞) + (h : ∀ x ∈ p, f x ∈ p) (h' : ⨆ μ, f.genEigenspace μ k = ⊤) : + p = ⨆ μ, p ⊓ f.genEigenspace μ k := by rw [← inf_iSup_genEigenspace h, h', inf_top_eq] end Submodule /-- In finite dimensions, if the generalized eigenspaces of a linear endomorphism span the whole space then the same is true of its restriction to any invariant submodule. -/ -theorem Module.End.unifEigenspace_restrict_eq_top +theorem Module.End.genEigenspace_restrict_eq_top {p : Submodule K V} {f : Module.End K V} [FiniteDimensional K V] {k : ℕ∞} - (h : ∀ x ∈ p, f x ∈ p) (h' : ⨆ μ, f.unifEigenspace μ k = ⊤) : - ⨆ μ, Module.End.unifEigenspace (LinearMap.restrict f h) μ k = ⊤ := by - have := congr_arg (Submodule.comap p.subtype) (Submodule.eq_iSup_inf_unifEigenspace k h h') + (h : ∀ x ∈ p, f x ∈ p) (h' : ⨆ μ, f.genEigenspace μ k = ⊤) : + ⨆ μ, Module.End.genEigenspace (LinearMap.restrict f h) μ k = ⊤ := by + have := congr_arg (Submodule.comap p.subtype) (Submodule.eq_iSup_inf_genEigenspace k h h') have h_inj : Function.Injective p.subtype := Subtype.coe_injective - simp_rw [Submodule.inf_unifEigenspace f p h, Submodule.comap_subtype_self, + simp_rw [Submodule.inf_genEigenspace f p h, Submodule.comap_subtype_self, ← Submodule.map_iSup, Submodule.comap_map_eq_of_injective h_inj] at this exact this.symm /-- In finite dimensions, if the generalized eigenspaces of a linear endomorphism span the whole space then the same is true of its restriction to any invariant submodule. -/ -@[deprecated Module.End.unifEigenspace_restrict_eq_top (since := "2024-10-11")] +@[deprecated Module.End.genEigenspace_restrict_eq_top (since := "2024-10-11")] theorem Module.End.iSup_genEigenspace_restrict_eq_top {p : Submodule K V} {f : Module.End K V} [FiniteDimensional K V] - (h : ∀ x ∈ p, f x ∈ p) (h' : ⨆ μ, ⨆ k, f.genEigenspace μ k = ⊤) : - ⨆ μ, ⨆ k, Module.End.genEigenspace (LinearMap.restrict f h) μ k = ⊤ := by + (h : ∀ x ∈ p, f x ∈ p) (h' : ⨆ μ, ⨆ k : ℕ, f.genEigenspace μ k = ⊤) : + ⨆ μ, ⨆ k : ℕ, Module.End.genEigenspace (LinearMap.restrict f h) μ k = ⊤ := by simp_rw [iSup_genEigenspace_eq] at h' ⊢ - apply Module.End.unifEigenspace_restrict_eq_top h h' + apply Module.End.genEigenspace_restrict_eq_top h h' diff --git a/Mathlib/LinearAlgebra/Eigenspace/Zero.lean b/Mathlib/LinearAlgebra/Eigenspace/Zero.lean index cbae5ab149388..21e08bdf55996 100644 --- a/Mathlib/LinearAlgebra/Eigenspace/Zero.lean +++ b/Mathlib/LinearAlgebra/Eigenspace/Zero.lean @@ -132,7 +132,7 @@ lemma finrank_maxGenEigenspace (φ : Module.End K M) : finrank K (φ.maxGenEigenspace 0) = natTrailingDegree (φ.charpoly) := by set V := φ.maxGenEigenspace 0 have hV : V = ⨆ (n : ℕ), ker (φ ^ n) := by - simp [V, ← Module.End.iSup_genEigenspace_eq, Module.End.genEigenspace_def] + simp [V, ← Module.End.iSup_genEigenspace_eq, Module.End.genEigenspace_nat] let W := ⨅ (n : ℕ), LinearMap.range (φ ^ n) have hVW : IsCompl V W := by rw [hV] From 8b008cbc2d4d67a9a8f01a265c8d984303f38049 Mon Sep 17 00:00:00 2001 From: YnirPaz Date: Sat, 26 Oct 2024 14:30:57 +0000 Subject: [PATCH 451/505] feat(SetTheory/Cardinal/Basic): out_lift_equiv (#16943) Prove `out_lift_equiv`. --- Mathlib/SetTheory/Cardinal/Basic.lean | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/Mathlib/SetTheory/Cardinal/Basic.lean b/Mathlib/SetTheory/Cardinal/Basic.lean index e44825db47401..d330b1ffc5407 100644 --- a/Mathlib/SetTheory/Cardinal/Basic.lean +++ b/Mathlib/SetTheory/Cardinal/Basic.lean @@ -266,6 +266,10 @@ theorem lift_uzero (a : Cardinal.{u}) : lift.{0} a = a := theorem lift_lift.{u_1} (a : Cardinal.{u_1}) : lift.{w} (lift.{v} a) = lift.{max v w} a := inductionOn a fun _ => (Equiv.ulift.trans <| Equiv.ulift.trans Equiv.ulift.symm).cardinal_eq +theorem out_lift_equiv (a : Cardinal.{u}) : Nonempty ((lift.{v} a).out ≃ a.out) := by + rw [← mk_out a, ← mk_uLift, mk_out] + exact ⟨outMkEquiv.trans Equiv.ulift⟩ + @[simp] lemma mk_preimage_down {s : Set α} : #(ULift.down.{v} ⁻¹' s) = lift.{v} (#s) := by rw [← mk_uLift, Cardinal.eq] From cc31026325589345000c05b2fae92c9a0b01ce05 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Sat, 26 Oct 2024 14:30:58 +0000 Subject: [PATCH 452/505] RFC/chore(Algebra/GCDMonoid): should we un-`@[simp]` normalize_apply? (#18006) The lemma `normalize_apply` feels to me like it exposes the implementation detail that `normalize` is defined in terms of `normUnit`. I think the `normalize` function itself is much more familiar so it doesn't really seem like a simplification to insert `normUnit` instead. Indeed there are quite a few `simp [-normalize_apply]` in Mathlib. Moreover, there were some `simp` lemmas about `normalize` that had to be given higher priority since e.g. `normalize (a * b) = normalize a * normalize b` but `normUnit (a * b)` does not necessarily equal `normUnit a * normUnit b`. A few downstream fixes are needed where a `@[simp]` lemma is proved for `normUnit` and not for `normalize`, otherwise everything goes very smoothly. --- Mathlib/Algebra/GCDMonoid/Basic.lean | 30 +++++++------------ Mathlib/Algebra/GCDMonoid/Nat.lean | 4 +-- Mathlib/Algebra/Polynomial/FieldDivision.lean | 7 +++-- Mathlib/RingTheory/DedekindDomain/Ideal.lean | 2 +- 4 files changed, 19 insertions(+), 24 deletions(-) diff --git a/Mathlib/Algebra/GCDMonoid/Basic.lean b/Mathlib/Algebra/GCDMonoid/Basic.lean index 1eb20199e470b..1eb6be1251a77 100644 --- a/Mathlib/Algebra/GCDMonoid/Basic.lean +++ b/Mathlib/Algebra/GCDMonoid/Basic.lean @@ -116,21 +116,16 @@ theorem normalize_associated_iff {x y : α} : Associated (normalize x) y ↔ Ass theorem Associates.mk_normalize (x : α) : Associates.mk (normalize x) = Associates.mk x := Associates.mk_eq_mk_iff_associated.2 (normalize_associated _) -@[simp] theorem normalize_apply (x : α) : normalize x = x * normUnit x := rfl --- Porting note (#10618): `simp` can prove this --- @[simp] theorem normalize_zero : normalize (0 : α) = 0 := normalize.map_zero --- Porting note (#10618): `simp` can prove this --- @[simp] theorem normalize_one : normalize (1 : α) = 1 := normalize.map_one -theorem normalize_coe_units (u : αˣ) : normalize (u : α) = 1 := by simp +theorem normalize_coe_units (u : αˣ) : normalize (u : α) = 1 := by simp [normalize_apply] theorem normalize_eq_zero {x : α} : normalize x = 0 ↔ x = 0 := ⟨fun hx => (associated_zero_iff_eq_zero x).1 <| hx ▸ associated_normalize _, by @@ -147,7 +142,8 @@ theorem normUnit_mul_normUnit (a : α) : normUnit (a * normUnit a) = 1 := by · rw [normUnit_zero, zero_mul, normUnit_zero] · rw [normUnit_mul h (Units.ne_zero _), normUnit_coe_units, mul_inv_eq_one] -theorem normalize_idem (x : α) : normalize (normalize x) = normalize x := by simp +@[simp] +theorem normalize_idem (x : α) : normalize (normalize x) = normalize x := by simp [normalize_apply] theorem normalize_eq_normalize {a b : α} (hab : a ∣ b) (hba : b ∣ a) : normalize a = normalize b := by @@ -173,11 +169,11 @@ theorem Associated.eq_of_normalized a = b := dvd_antisymm_of_normalize_eq ha hb h.dvd h.dvd' ---can be proven by simp +@[simp] theorem dvd_normalize_iff {a b : α} : a ∣ normalize b ↔ a ∣ b := Units.dvd_mul_right ---can be proven by simp +@[simp] theorem normalize_dvd_iff {a b : α} : normalize a ∣ b ↔ a ∣ b := Units.mul_right_dvd @@ -216,8 +212,7 @@ theorem out_dvd_iff (a : α) (b : Associates α) : b.out ∣ a ↔ b ≤ Associa theorem out_top : (⊤ : Associates α).out = 0 := normalize_zero --- Porting note: lower priority to avoid linter complaints about simp-normal form -@[simp 1100] +@[simp] theorem normalize_out (a : Associates α) : normalize a.out = a.out := Quotient.inductionOn a normalize_idem @@ -287,8 +282,7 @@ theorem gcd_isUnit_iff_isRelPrime [GCDMonoid α] {a b : α} : IsUnit (gcd a b) ↔ IsRelPrime a b := ⟨fun h _ ha hb ↦ isUnit_of_dvd_unit (dvd_gcd ha hb) h, (· (gcd_dvd_left a b) (gcd_dvd_right a b))⟩ --- Porting note: lower priority to avoid linter complaints about simp-normal form -@[simp 1100] +@[simp] theorem normalize_gcd [NormalizedGCDMonoid α] : ∀ a b : α, normalize (gcd a b) = gcd a b := NormalizedGCDMonoid.normalize_gcd @@ -399,7 +393,7 @@ theorem gcd_mul_left [NormalizedGCDMonoid α] (a b c : α) : gcd (a * b) (a * c) = normalize a * gcd b c := (by_cases (by rintro rfl; simp only [zero_mul, gcd_zero_left, normalize_zero])) fun ha : a ≠ 0 => - suffices gcd (a * b) (a * c) = normalize (a * gcd b c) by simpa [- normalize_apply] + suffices gcd (a * b) (a * c) = normalize (a * gcd b c) by simpa let ⟨d, eq⟩ := dvd_gcd (dvd_mul_right a b) (dvd_mul_right a c) gcd_eq_normalize (eq.symm ▸ mul_dvd_mul_left a @@ -689,8 +683,7 @@ theorem lcm_eq_zero_iff [GCDMonoid α] (a b : α) : lcm a b = 0 ↔ a = 0 ∨ b rwa [← mul_eq_zero, ← associated_zero_iff_eq_zero]) (by rintro (rfl | rfl) <;> [apply lcm_zero_left; apply lcm_zero_right]) --- Porting note: lower priority to avoid linter complaints about simp-normal form -@[simp 1100] +@[simp] theorem normalize_lcm [NormalizedGCDMonoid α] (a b : α) : normalize (lcm a b) = lcm a b := NormalizedGCDMonoid.normalize_lcm a b @@ -766,7 +759,7 @@ theorem lcm_mul_left [NormalizedGCDMonoid α] (a b c : α) : lcm (a * b) (a * c) = normalize a * lcm b c := (by_cases (by rintro rfl; simp only [zero_mul, lcm_zero_left, normalize_zero])) fun ha : a ≠ 0 => - suffices lcm (a * b) (a * c) = normalize (a * lcm b c) by simpa [- normalize_apply] + suffices lcm (a * b) (a * c) = normalize (a * lcm b c) by simpa have : a ∣ lcm (a * b) (a * c) := (dvd_mul_right _ _).trans (dvd_lcm_left _ _) let ⟨_, eq⟩ := this lcm_eq_normalize @@ -866,8 +859,7 @@ instance subsingleton_normalizedGCDMonoid_of_unique_units : Subsingleton (Normal theorem normUnit_eq_one (x : α) : normUnit x = 1 := rfl --- Porting note (#10618): `simp` can prove this --- @[simp] +@[simp] theorem normalize_eq (x : α) : normalize x = x := mul_one x diff --git a/Mathlib/Algebra/GCDMonoid/Nat.lean b/Mathlib/Algebra/GCDMonoid/Nat.lean index 2824c117566c0..a36ea2e3acbc3 100644 --- a/Mathlib/Algebra/GCDMonoid/Nat.lean +++ b/Mathlib/Algebra/GCDMonoid/Nat.lean @@ -79,7 +79,7 @@ theorem normalize_coe_nat (n : ℕ) : normalize (n : ℤ) = n := theorem abs_eq_normalize (z : ℤ) : |z| = normalize z := by cases le_total 0 z <;> - simp [-normalize_apply, abs_of_nonneg, abs_of_nonpos, normalize_of_nonneg, normalize_of_nonpos, *] + simp [abs_of_nonneg, abs_of_nonpos, normalize_of_nonneg, normalize_of_nonpos, *] theorem nonneg_of_normalize_eq_self {z : ℤ} (hz : normalize z = z) : 0 ≤ z := by by_cases h : 0 ≤ z @@ -147,7 +147,7 @@ def associatesIntEquivNat : Associates ℤ ≃ ℕ := by refine ⟨(·.out.natAbs), (Associates.mk ·), ?_, fun n ↦ ?_⟩ · refine Associates.forall_associated.2 fun a ↦ ?_ refine Associates.mk_eq_mk_iff_associated.2 <| Associated.symm <| ⟨normUnit a, ?_⟩ - simp [Int.abs_eq_normalize] + simp [Int.abs_eq_normalize, normalize_apply] · dsimp only [Associates.out_mk] rw [← Int.abs_eq_normalize, Int.natAbs_abs, Int.natAbs_ofNat] diff --git a/Mathlib/Algebra/Polynomial/FieldDivision.lean b/Mathlib/Algebra/Polynomial/FieldDivision.lean index f59c7998f4242..b0f395c366efb 100644 --- a/Mathlib/Algebra/Polynomial/FieldDivision.lean +++ b/Mathlib/Algebra/Polynomial/FieldDivision.lean @@ -195,8 +195,9 @@ instance instNormalizationMonoid : NormalizationMonoid R[X] where theorem coe_normUnit {p : R[X]} : (normUnit p : R[X]) = C ↑(normUnit p.leadingCoeff) := by simp [normUnit] +@[simp] theorem leadingCoeff_normalize (p : R[X]) : - leadingCoeff (normalize p) = normalize (leadingCoeff p) := by simp + leadingCoeff (normalize p) = normalize (leadingCoeff p) := by simp [normalize_apply] theorem Monic.normalize_eq_self {p : R[X]} (hp : p.Monic) : normalize p = p := by simp only [Polynomial.coe_normUnit, normalize_apply, hp.leadingCoeff, normUnit_one, @@ -536,7 +537,9 @@ theorem map_dvd_map' [Field k] (f : R →+* k) {x y : R[X]} : x.map f ∣ y.map leadingCoeff_map, ← map_inv₀ f, ← map_C, ← Polynomial.map_mul, map_dvd_map _ f.injective (monic_mul_leadingCoeff_inv H)] -theorem degree_normalize [DecidableEq R] : degree (normalize p) = degree p := by simp +@[simp] +theorem degree_normalize [DecidableEq R] : degree (normalize p) = degree p := by + simp [normalize_apply] theorem prime_of_degree_eq_one (hp1 : degree p = 1) : Prime p := by classical diff --git a/Mathlib/RingTheory/DedekindDomain/Ideal.lean b/Mathlib/RingTheory/DedekindDomain/Ideal.lean index 3fae07bf9f31e..b4067b230ce8c 100644 --- a/Mathlib/RingTheory/DedekindDomain/Ideal.lean +++ b/Mathlib/RingTheory/DedekindDomain/Ideal.lean @@ -1429,7 +1429,7 @@ theorem count_span_normalizedFactors_eq_of_normUnit {r X : R} (hr : r ≠ 0) (hX₁ : normUnit X = 1) (hX : Prime X) : Multiset.count (Ideal.span {X} : Ideal R) (normalizedFactors (Ideal.span {r})) = Multiset.count X (normalizedFactors r) := by - simpa [hX₁] using count_span_normalizedFactors_eq hr hX + simpa [hX₁, normalize_apply] using count_span_normalizedFactors_eq hr hX end NormalizationMonoid From 2e0b67b8e495c0c8b7162686b6b639101a7afef8 Mon Sep 17 00:00:00 2001 From: Jakob von Raumer Date: Sat, 26 Oct 2024 14:30:59 +0000 Subject: [PATCH 453/505] feat(CategoryTheory/Grothendieck): `Grothendieck.pre` for Grothendieck construction base change (#18044) Needed in order to use the Grothendieck construction as a way to express dependent double colimits. This also makes the following changes: * Improves `Grothendieck.comp_base` to trigger without using `erw` (by deriving it manually instead of by `@[simps]` * Generalizes `IsConnected.zigzag_obj_of_zigzag` from functors to prefunctors since proving functoriality at use sites makes no sense (since we throw away the data anyway and just keep a disjunction of a `Nonempty` of morphism types. --- Mathlib/CategoryTheory/Grothendieck.lean | 32 +++++++++++++----- Mathlib/CategoryTheory/IsConnected.lean | 9 ++++- Mathlib/CategoryTheory/Limits/Final.lean | 43 ++++++++++++++++++++++++ 3 files changed, 74 insertions(+), 10 deletions(-) diff --git a/Mathlib/CategoryTheory/Grothendieck.lean b/Mathlib/CategoryTheory/Grothendieck.lean index 5a361423ddbc9..a6af93f1d5667 100644 --- a/Mathlib/CategoryTheory/Grothendieck.lean +++ b/Mathlib/CategoryTheory/Grothendieck.lean @@ -94,7 +94,6 @@ instance (X : Grothendieck F) : Inhabited (Hom X X) := /-- Composition of morphisms in the Grothendieck category. -/ -@[simps] def comp {X Y Z : Grothendieck F} (f : Hom X Y) (g : Hom Y Z) : Hom X Z where base := f.base ≫ g.base fiber := @@ -108,15 +107,15 @@ instance : Category (Grothendieck F) where comp := @fun _ _ _ f g => Grothendieck.comp f g comp_id := @fun X Y f => by dsimp; ext - · simp - · dsimp + · simp [comp] + · dsimp [comp] rw [← NatIso.naturality_2 (eqToIso (F.map_id Y.base)) f.fiber] simp - id_comp := @fun X Y f => by dsimp; ext <;> simp + id_comp := @fun X Y f => by dsimp; ext <;> simp [comp] assoc := @fun W X Y Z f g h => by dsimp; ext - · simp - · dsimp + · simp [comp] + · dsimp [comp] rw [← NatIso.naturality_2 (eqToIso (F.map_comp _ _)) f.fiber] simp @@ -126,11 +125,16 @@ theorem id_fiber' (X : Grothendieck F) : id_fiber X @[simp] -theorem comp_fiber' {X Y Z : Grothendieck F} (f : X ⟶ Y) (g : Y ⟶ Z) : +theorem comp_base {X Y Z : Grothendieck F} (f : X ⟶ Y) (g : Y ⟶ Z) : + (f ≫ g).base = f.base ≫ g.base := + rfl + +@[simp] +theorem comp_fiber {X Y Z : Grothendieck F} (f : X ⟶ Y) (g : Y ⟶ Z) : Hom.fiber (f ≫ g) = eqToHom (by erw [Functor.map_comp, Functor.comp_obj]) ≫ (F.map g.base).map f.fiber ≫ g.fiber := - comp_fiber f g + rfl theorem congr {X Y : Grothendieck F} {f g : X ⟶ Y} (h : f = g) : @@ -170,7 +174,7 @@ def map (α : F ⟶ G) : Grothendieck F ⥤ Grothendieck G where map_comp {X Y Z} f g := by dsimp congr 1 - simp only [comp_fiber' f g, ← Category.assoc, Functor.map_comp, eqToHom_map] + simp only [comp_fiber f g, ← Category.assoc, Functor.map_comp, eqToHom_map] congr 1 simp only [Cat.eqToHom_app, Cat.comp_obj, eqToHom_trans, eqToHom_map, Category.assoc] erw [Functor.congr_hom (α.naturality g.base).symm f.fiber] @@ -292,6 +296,16 @@ def grothendieckTypeToCat : Grothendieck (G ⋙ typeToCat) ≌ G.Elements where simp rfl +variable (F) in +/-- Applying a functor `G : D ⥤ C` to the base of the Grothendieck construction induces a functor +`Grothendieck (G ⋙ F) ⥤ Grothendieck F`. -/ +@[simps] +def pre (G : D ⥤ C) : Grothendieck (G ⋙ F) ⥤ Grothendieck F where + obj X := ⟨G.obj X.base, X.fiber⟩ + map f := ⟨G.map f.base, f.fiber⟩ + map_id X := Grothendieck.ext _ _ (G.map_id _) (by simp) + map_comp f g := Grothendieck.ext _ _ (G.map_comp _ _) (by simp) + end Grothendieck end CategoryTheory diff --git a/Mathlib/CategoryTheory/IsConnected.lean b/Mathlib/CategoryTheory/IsConnected.lean index c6b21cf4cbce6..fc0e0c664675e 100644 --- a/Mathlib/CategoryTheory/IsConnected.lean +++ b/Mathlib/CategoryTheory/IsConnected.lean @@ -324,12 +324,19 @@ def Zigzag.setoid (J : Type u₂) [Category.{v₁} J] : Setoid J where r := Zigzag iseqv := zigzag_equivalence +/-- If there is a zigzag from `j₁` to `j₂`, then there is a zigzag from `F j₁` to +`F j₂` as long as `F` is a prefunctor. +-/ +theorem zigzag_prefunctor_obj_of_zigzag (F : J ⥤q K) {j₁ j₂ : J} (h : Zigzag j₁ j₂) : + Zigzag (F.obj j₁) (F.obj j₂) := + h.lift _ fun _ _ => Or.imp (Nonempty.map fun f => F.map f) (Nonempty.map fun f => F.map f) + /-- If there is a zigzag from `j₁` to `j₂`, then there is a zigzag from `F j₁` to `F j₂` as long as `F` is a functor. -/ theorem zigzag_obj_of_zigzag (F : J ⥤ K) {j₁ j₂ : J} (h : Zigzag j₁ j₂) : Zigzag (F.obj j₁) (F.obj j₂) := - h.lift _ fun _ _ => Or.imp (Nonempty.map fun f => F.map f) (Nonempty.map fun f => F.map f) + zigzag_prefunctor_obj_of_zigzag F.toPrefunctor h /-- A Zag in a discrete category entails an equality of its extremities -/ lemma eq_of_zag (X) {a b : Discrete X} (h : Zag a b) : a.as = b.as := diff --git a/Mathlib/CategoryTheory/Limits/Final.lean b/Mathlib/CategoryTheory/Limits/Final.lean index 6390aaa372617..2cd80eb36a4f4 100644 --- a/Mathlib/CategoryTheory/Limits/Final.lean +++ b/Mathlib/CategoryTheory/Limits/Final.lean @@ -10,6 +10,7 @@ import Mathlib.CategoryTheory.Limits.Shapes.Types import Mathlib.CategoryTheory.Filtered.Basic import Mathlib.CategoryTheory.Limits.Yoneda import Mathlib.CategoryTheory.PUnit +import Mathlib.CategoryTheory.Grothendieck /-! # Final and initial functors @@ -848,4 +849,46 @@ instance CostructuredArrow.initial_pre (T : C ⥤ D) [Initial T] (S : D ⥤ E) ( end +section Grothendieck + +variable {C : Type u₁} [Category.{v₁} C] +variable {D : Type u₂} [Category.{v₂} D] +variable (F : D ⥤ Cat) (G : C ⥤ D) + +open Functor + +/-- A prefunctor mapping structured arrows on `G` to structured arrows on `pre F G` with their +action on fibers being the identity. -/ +def Grothendieck.structuredArrowToStructuredArrowPre (d : D) (f : F.obj d) : + StructuredArrow d G ⥤q StructuredArrow ⟨d, f⟩ (pre F G) where + obj := fun X => StructuredArrow.mk (Y := ⟨X.right, (F.map X.hom).obj f⟩) + (Grothendieck.Hom.mk (by exact X.hom) (by dsimp; exact 𝟙 _)) + map := fun g => StructuredArrow.homMk + (Grothendieck.Hom.mk (by exact g.right) + (eqToHom (by dsimp; rw [← StructuredArrow.w g, map_comp, Cat.comp_obj]))) + (by simp only [StructuredArrow.mk_right] + apply Grothendieck.ext <;> simp) + +instance Grothendieck.final_pre [hG : Final G] : (Grothendieck.pre F G).Final := by + constructor + rintro ⟨d, f⟩ + let ⟨u, c, g⟩ : Nonempty (StructuredArrow d G) := inferInstance + letI : Nonempty (StructuredArrow ⟨d, f⟩ (pre F G)) := + ⟨u, ⟨c, (F.map g).obj f⟩, ⟨(by exact g), (by exact 𝟙 _)⟩⟩ + apply zigzag_isConnected + rintro ⟨⟨⟨⟩⟩, ⟨bi, fi⟩, ⟨gbi, gfi⟩⟩ ⟨⟨⟨⟩⟩, ⟨bj, fj⟩, ⟨gbj, gfj⟩⟩ + dsimp at fj fi gfi gbi gbj gfj + apply Zigzag.trans (j₂ := StructuredArrow.mk (Y := ⟨bi, ((F.map gbi).obj f)⟩) + (Grothendieck.Hom.mk gbi (𝟙 _))) + (.of_zag (.inr ⟨StructuredArrow.homMk (Grothendieck.Hom.mk (by dsimp; exact 𝟙 _) + (eqToHom (by simp) ≫ gfi)) (by apply Grothendieck.ext <;> simp)⟩)) + refine Zigzag.trans (j₂ := StructuredArrow.mk (Y := ⟨bj, ((F.map gbj).obj f)⟩) + (Grothendieck.Hom.mk gbj (𝟙 _))) ?_ + (.of_zag (.inl ⟨StructuredArrow.homMk (Grothendieck.Hom.mk (by dsimp; exact 𝟙 _) + (eqToHom (by simp) ≫ gfj)) (by apply Grothendieck.ext <;> simp)⟩)) + exact zigzag_prefunctor_obj_of_zigzag (Grothendieck.structuredArrowToStructuredArrowPre F G d f) + (isPreconnected_zigzag (.mk gbi) (.mk gbj)) + +end Grothendieck + end CategoryTheory From 85f9ce7d1f62a60e4bf85bf1fb612f88a8c9c65d Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 26 Oct 2024 15:15:15 +0000 Subject: [PATCH 454/505] feat: more lemmas about `Fin`-indexed products of finsets (#17499) Also use the new sexy `Finset.filter` notation --- Mathlib/Data/Fin/Tuple/Finset.lean | 39 +++++++++++++++++++++--------- 1 file changed, 27 insertions(+), 12 deletions(-) diff --git a/Mathlib/Data/Fin/Tuple/Finset.lean b/Mathlib/Data/Fin/Tuple/Finset.lean index d376279817e25..7c012e3737704 100644 --- a/Mathlib/Data/Fin/Tuple/Finset.lean +++ b/Mathlib/Data/Fin/Tuple/Finset.lean @@ -57,33 +57,48 @@ namespace Finset variable {n : ℕ} {α : Fin (n + 1) → Type*} {p : Fin (n + 1)} (S : ∀ i, Finset (α i)) lemma map_consEquiv_filter_piFinset (P : (∀ i, α (succ i)) → Prop) [DecidablePred P] : - ((piFinset S).filter fun r ↦ P <| tail r).map (consEquiv α).symm.toEmbedding = - S 0 ×ˢ (piFinset fun x ↦ S <| succ x).filter P := by + {r ∈ piFinset S | P (tail r)}.map (consEquiv α).symm.toEmbedding = + S 0 ×ˢ {r ∈ piFinset (tail S) | P r} := by unfold tail; ext; simp [Fin.forall_iff_succ, and_assoc] lemma map_snocEquiv_filter_piFinset (P : (∀ i, α (castSucc i)) → Prop) [DecidablePred P] : - ((piFinset S).filter fun r ↦ P <| init r).map (snocEquiv α).symm.toEmbedding = - S (last _) ×ˢ (piFinset <| init S).filter P := by + {r ∈ piFinset S | P (init r)}.map (snocEquiv α).symm.toEmbedding = + S (last _) ×ˢ {r ∈ piFinset (init S) | P r} := by unfold init; ext; simp [Fin.forall_iff_castSucc, and_assoc] lemma map_insertNthEquiv_filter_piFinset (P : (∀ i, α (p.succAbove i)) → Prop) [DecidablePred P] : - ((piFinset S).filter fun r ↦ P <| p.removeNth r).map (p.insertNthEquiv α).symm.toEmbedding = - S p ×ˢ (piFinset <| p.removeNth S).filter P := by + {r ∈ piFinset S | P (p.removeNth r)}.map (p.insertNthEquiv α).symm.toEmbedding = + S p ×ˢ {r ∈ piFinset (p.removeNth S) | P r} := by unfold removeNth; ext; simp [Fin.forall_iff_succAbove p, and_assoc] +lemma filter_piFinset_eq_map_consEquiv (P : (∀ i, α (succ i)) → Prop) [DecidablePred P] : + {r ∈ piFinset S | P (tail r)} = + (S 0 ×ˢ {r ∈ piFinset (tail S) | P r}).map (consEquiv α).toEmbedding := by + simp [← map_consEquiv_filter_piFinset, map_map] + +lemma filter_piFinset_eq_map_snocEquiv (P : (∀ i, α (castSucc i)) → Prop) [DecidablePred P] : + {r ∈ piFinset S | P (init r)} = + (S (last _) ×ˢ {r ∈ piFinset (init S) | P r}).map (snocEquiv α).toEmbedding := by + simp [← map_snocEquiv_filter_piFinset, map_map] + +lemma filter_piFinset_eq_map_insertNthEquiv (P : (∀ i, α (p.succAbove i)) → Prop) + [DecidablePred P] : + {r ∈ piFinset S | P (p.removeNth r)} = + (S p ×ˢ {r ∈ piFinset (p.removeNth S) | P r}).map (p.insertNthEquiv α).toEmbedding := by + simp [← map_insertNthEquiv_filter_piFinset, map_map] + lemma card_consEquiv_filter_piFinset (P : (∀ i, α (succ i)) → Prop) [DecidablePred P] : - ((piFinset S).filter fun r ↦ P <| tail r).card = - (S 0).card * ((piFinset fun x ↦ S <| succ x).filter P).card := by + {r ∈ piFinset S | P (tail r)}.card = (S 0).card * {r ∈ piFinset (tail S) | P r}.card := by rw [← card_product, ← map_consEquiv_filter_piFinset, card_map] lemma card_snocEquiv_filter_piFinset (P : (∀ i, α (castSucc i)) → Prop) [DecidablePred P] : - ((piFinset S).filter fun r ↦ P <| init r).card = - (S (last _)).card * ((piFinset <| init S).filter P).card := by + {r ∈ piFinset S | P (init r)}.card = + (S (last _)).card * {r ∈ piFinset (init S) | P r}.card := by rw [← card_product, ← map_snocEquiv_filter_piFinset, card_map] lemma card_insertNthEquiv_filter_piFinset (P : (∀ i, α (p.succAbove i)) → Prop) [DecidablePred P] : - ((piFinset S).filter fun r ↦ P <| p.removeNth r).card = - (S p).card * ((piFinset <| p.removeNth S).filter P).card := by + {r ∈ piFinset S | P (p.removeNth r)}.card = + (S p).card * {r ∈ piFinset (p.removeNth S) | P r}.card := by rw [← card_product, ← map_insertNthEquiv_filter_piFinset, card_map] end Finset From b3092c0317cc2a3a14a03b1bbabc311e4cb3d54e Mon Sep 17 00:00:00 2001 From: Yoh Tanimoto Date: Sat, 26 Oct 2024 15:15:17 +0000 Subject: [PATCH 455/505] feat(Topology/UnitInterval): add unitInterval as Submonoid (#17844) add the definition of `UnitInterval` as Submonoid of `Real`. motivation: This enables the use of `prod_mem`. suggested in #12266 Co-authored-by: Johan Commelin --- Mathlib/Topology/UnitInterval.lean | 14 ++++++++++++++ 1 file changed, 14 insertions(+) diff --git a/Mathlib/Topology/UnitInterval.lean b/Mathlib/Topology/UnitInterval.lean index 55c6405dda1b6..40404390094c0 100644 --- a/Mathlib/Topology/UnitInterval.lean +++ b/Mathlib/Topology/UnitInterval.lean @@ -209,6 +209,20 @@ theorem mul_pos_mem_iff {a t : ℝ} (ha : 0 < a) : a * t ∈ I ↔ t ∈ Set.Icc theorem two_mul_sub_one_mem_iff {t : ℝ} : 2 * t - 1 ∈ I ↔ t ∈ Set.Icc (1 / 2 : ℝ) 1 := by constructor <;> rintro ⟨h₁, h₂⟩ <;> constructor <;> linarith +/-- The unit interval as a submonoid of ℝ. -/ +def submonoid : Submonoid ℝ where + carrier := unitInterval + one_mem' := unitInterval.one_mem + mul_mem' := unitInterval.mul_mem + +@[simp] theorem coe_unitIntervalSubmonoid : submonoid = unitInterval := rfl +@[simp] theorem mem_unitIntervalSubmonoid {x} : x ∈ submonoid ↔ x ∈ unitInterval := + Iff.rfl + +protected theorem prod_mem {ι : Type*} {t : Finset ι} {f : ι → ℝ} + (h : ∀ c ∈ t, f c ∈ unitInterval) : + ∏ c ∈ t, f c ∈ unitInterval := _root_.prod_mem (S := unitInterval.submonoid) h + instance : LinearOrderedCommMonoidWithZero I where zero_mul i := zero_mul i mul_zero i := mul_zero i From 8cb7d6fc4d7757deaadb42fabaf8c2a1aaffa11a Mon Sep 17 00:00:00 2001 From: Yakov Pechersky Date: Sat, 26 Oct 2024 15:46:19 +0000 Subject: [PATCH 456/505] feat(Ideal/IsPrincipalPowQuotient): R/I equiv I^n/I^(n+1) (#15426) Co-authored-by: Yakov Pechersky --- Mathlib.lean | 1 + .../Ideal/IsPrincipalPowQuotient.lean | 86 +++++++++++++++++++ .../RingTheory/Ideal/QuotientOperations.lean | 44 ++++++++++ 3 files changed, 131 insertions(+) create mode 100644 Mathlib/RingTheory/Ideal/IsPrincipalPowQuotient.lean diff --git a/Mathlib.lean b/Mathlib.lean index 802e7e7de1d6b..05c28e27f3830 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4048,6 +4048,7 @@ import Mathlib.RingTheory.Ideal.Cotangent import Mathlib.RingTheory.Ideal.IdempotentFG import Mathlib.RingTheory.Ideal.IsPrimary import Mathlib.RingTheory.Ideal.IsPrincipal +import Mathlib.RingTheory.Ideal.IsPrincipalPowQuotient import Mathlib.RingTheory.Ideal.Maps import Mathlib.RingTheory.Ideal.MinimalPrime import Mathlib.RingTheory.Ideal.Norm diff --git a/Mathlib/RingTheory/Ideal/IsPrincipalPowQuotient.lean b/Mathlib/RingTheory/Ideal/IsPrincipalPowQuotient.lean new file mode 100644 index 0000000000000..2255eaacad2a1 --- /dev/null +++ b/Mathlib/RingTheory/Ideal/IsPrincipalPowQuotient.lean @@ -0,0 +1,86 @@ +/- +Copyright (c) 2024 Yakov Pechersky. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yakov Pechersky +-/ +import Mathlib.LinearAlgebra.Isomorphisms +import Mathlib.RingTheory.Ideal.Operations +import Mathlib.RingTheory.Ideal.Quotient + +/-! +# Quotients of powers of principal ideals + +This file deals with taking quotients of powers of principal ideals. + +## Main definitions and results + +* `Ideal.quotEquivPowQuotPowSucc`: for a principal ideal `I`, `R ⧸ I ≃ₗ[R] I ^ n ⧸ I ^ (n + 1)` + +## Implementation details + +At site of usage, calling `LinearEquiv.toEquiv` can cause timeouts in the search for a complex +synthesis like `Module 𝒪[K] 𝓀[k]`, so the plain equiv versions are provided. + +These equivs are defined here as opposed to in the quotients file since they cannot be +formed as ring equivs. + +-/ + + +namespace Ideal + +section IsPrincipal + +variable {R : Type*} [CommRing R] [IsDomain R] {I : Ideal R} + +/-- For a principal ideal `I`, `R ⧸ I ≃ₗ[R] I ^ n ⧸ I ^ (n + 1)`. To convert into a form +that uses the ideal of `R ⧸ I ^ (n + 1)`, compose with +`Ideal.powQuotPowSuccLinearEquivMapMkPowSuccPow`. -/ +noncomputable +def quotEquivPowQuotPowSucc (h : I.IsPrincipal) (h': I ≠ ⊥) (n : ℕ) : + (R ⧸ I) ≃ₗ[R] (I ^ n : Ideal R) ⧸ (I • ⊤ : Submodule R (I ^ n : Ideal R)) := by + let f : (I ^ n : Ideal R) →ₗ[R] (I ^ n : Ideal R) ⧸ (I • ⊤ : Submodule R (I ^ n : Ideal R)) := + Submodule.mkQ _ + let ϖ := h.principal'.choose + have hI : I = Ideal.span {ϖ} := h.principal'.choose_spec + have hϖ : ϖ ^ n ∈ I ^ n := hI ▸ (Ideal.pow_mem_pow (Ideal.mem_span_singleton_self _) n) + let g : R →ₗ[R] (I ^ n : Ideal R) := (LinearMap.mulRight R ϖ^n).codRestrict _ fun x ↦ by + simp only [LinearMap.pow_mulRight, LinearMap.mulRight_apply, Ideal.submodule_span_eq] + -- TODO: change argument of Ideal.pow_mem_of_mem + exact Ideal.mul_mem_left _ _ hϖ + have : I = LinearMap.ker (f.comp g) := by + ext x + simp only [LinearMap.codRestrict, LinearMap.pow_mulRight, LinearMap.mulRight_apply, + LinearMap.mem_ker, LinearMap.coe_comp, LinearMap.coe_mk, AddHom.coe_mk, Function.comp_apply, + Submodule.mkQ_apply, Submodule.Quotient.mk_eq_zero, Submodule.mem_smul_top_iff, smul_eq_mul, + f, g] + constructor <;> intro hx + · exact Submodule.mul_mem_mul hx hϖ + · rw [← pow_succ', hI, Ideal.span_singleton_pow, Ideal.mem_span_singleton] at hx + obtain ⟨y, hy⟩ := hx + rw [mul_comm, pow_succ, mul_assoc, mul_right_inj' (pow_ne_zero _ _)] at hy + · rw [hI, Ideal.mem_span_singleton] + exact ⟨y, hy⟩ + · contrapose! h' + rw [hI, h', Ideal.span_singleton_eq_bot] + let e : (R ⧸ I) ≃ₗ[R] R ⧸ (LinearMap.ker (f.comp g)) := + Submodule.quotEquivOfEq I (LinearMap.ker (f ∘ₗ g)) this + refine e.trans ((f.comp g).quotKerEquivOfSurjective ?_) + refine (Submodule.mkQ_surjective _).comp ?_ + rintro ⟨x, hx⟩ + rw [hI, Ideal.span_singleton_pow, Ideal.mem_span_singleton] at hx + refine hx.imp ?_ + simp [g, LinearMap.codRestrict, eq_comm, mul_comm] + +/-- For a principal ideal `I`, `R ⧸ I ≃ I ^ n ⧸ I ^ (n + 1)`. Supplied as a plain equiv to bypass +typeclass synthesis issues on complex `Module` goals. To convert into a form +that uses the ideal of `R ⧸ I ^ (n + 1)`, compose with +`Ideal.powQuotPowSuccEquivMapMkPowSuccPow`. -/ +noncomputable +def quotEquivPowQuotPowSuccEquiv (h : I.IsPrincipal) (h': I ≠ ⊥) (n : ℕ) : + (R ⧸ I) ≃ (I ^ n : Ideal R) ⧸ (I • ⊤ : Submodule R (I ^ n : Ideal R)) := + quotEquivPowQuotPowSucc h h' n + +end IsPrincipal + +end Ideal diff --git a/Mathlib/RingTheory/Ideal/QuotientOperations.lean b/Mathlib/RingTheory/Ideal/QuotientOperations.lean index efc97f56b8993..f189e59b5ea4e 100644 --- a/Mathlib/RingTheory/Ideal/QuotientOperations.lean +++ b/Mathlib/RingTheory/Ideal/QuotientOperations.lean @@ -938,3 +938,47 @@ theorem quotQuotEquivQuotOfLE_symm_comp_mkₐ (h : I ≤ J) : end AlgebraQuotient end DoubleQuot + +namespace Ideal + +section PowQuot + +variable {R : Type*} [CommRing R] (I : Ideal R) (n : ℕ) + +/-- `I ^ n ⧸ I ^ (n + 1)` can be viewed as a quotient module and as ideal of `R ⧸ I ^ (n + 1)`. +This definition gives the `R`-linear equivalence between the two. -/ +noncomputable +def powQuotPowSuccLinearEquivMapMkPowSuccPow : + ((I ^ n : Ideal R) ⧸ (I • ⊤ : Submodule R (I ^ n : Ideal R))) ≃ₗ[R] + Ideal.map (Ideal.Quotient.mk (I ^ (n + 1))) (I ^ n) := by + refine { LinearMap.codRestrict + (Submodule.restrictScalars _ (Ideal.map (Ideal.Quotient.mk (I ^ (n + 1))) (I ^ n))) + (Submodule.mapQ (I • ⊤) (I ^ (n + 1)) (Submodule.subtype (I ^ n)) ?_) ?_, + Equiv.ofBijective _ ⟨?_, ?_⟩ with } + · intro + simp [Submodule.mem_smul_top_iff, pow_succ'] + · intro x + obtain ⟨⟨y, hy⟩, rfl⟩ := Submodule.Quotient.mk_surjective _ x + simp [Ideal.mem_sup_left hy] + · intro a b + obtain ⟨⟨x, hx⟩, rfl⟩ := Submodule.Quotient.mk_surjective _ a + obtain ⟨⟨y, hy⟩, rfl⟩ := Submodule.Quotient.mk_surjective _ b + simp [Ideal.Quotient.eq, Submodule.Quotient.eq, Submodule.mem_smul_top_iff, pow_succ'] + · intro ⟨x, hx⟩ + rw [Ideal.mem_map_iff_of_surjective _ Ideal.Quotient.mk_surjective] at hx + obtain ⟨y, hy, rfl⟩ := hx + refine ⟨Submodule.Quotient.mk ⟨y, hy⟩, ?_⟩ + simp + +/-- `I ^ n ⧸ I ^ (n + 1)` can be viewed as a quotient module and as ideal of `R ⧸ I ^ (n + 1)`. +This definition gives the equivalence between the two, instead of the `R`-linear equivalence, +to bypass typeclass synthesis issues on complex `Module` goals. -/ +noncomputable +def powQuotPowSuccEquivMapMkPowSuccPow : + ((I ^ n : Ideal R) ⧸ (I • ⊤ : Submodule R (I ^ n : Ideal R))) ≃ + Ideal.map (Ideal.Quotient.mk (I ^ (n + 1))) (I ^ n) := + powQuotPowSuccLinearEquivMapMkPowSuccPow I n + +end PowQuot + +end Ideal From bf23ca5445dd7509615c489b06dbfc93546c1276 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 26 Oct 2024 15:58:48 +0000 Subject: [PATCH 457/505] feat: virtually nilpotent groups (#17628) Show that nilpotency is preserved under group isomorphisms and define virtually nilpotent subgroups. From GrowthInGroups --- Mathlib/GroupTheory/Nilpotent.lean | 42 +++++++++++++++++++++++++----- 1 file changed, 36 insertions(+), 6 deletions(-) diff --git a/Mathlib/GroupTheory/Nilpotent.lean b/Mathlib/GroupTheory/Nilpotent.lean index 14008d09636c0..9d9b75bfeff2d 100644 --- a/Mathlib/GroupTheory/Nilpotent.lean +++ b/Mathlib/GroupTheory/Nilpotent.lean @@ -144,25 +144,55 @@ theorem upperCentralSeries_one : upperCentralSeries G 1 = center G := by Subgroup.mem_center_iff, mem_mk, mem_bot, Set.mem_setOf_eq] exact forall_congr' fun y => by rw [mul_inv_eq_one, mul_inv_eq_iff_eq_mul, eq_comm] +variable {G} + /-- The `n+1`st term of the upper central series `H i` has underlying set equal to the `x` such that `⁅x,G⁆ ⊆ H n`-/ -theorem mem_upperCentralSeries_succ_iff (n : ℕ) (x : G) : +theorem mem_upperCentralSeries_succ_iff {n : ℕ} {x : G} : x ∈ upperCentralSeries G (n + 1) ↔ ∀ y : G, x * y * x⁻¹ * y⁻¹ ∈ upperCentralSeries G n := Iff.rfl +@[simp] lemma comap_upperCentralSeries {H : Type*} [Group H] (e : H ≃* G) : + ∀ n, (upperCentralSeries G n).comap e = upperCentralSeries H n + | 0 => by simpa [MonoidHom.ker_eq_bot_iff] using e.injective + | n + 1 => by + ext + simp [mem_upperCentralSeries_succ_iff, ← comap_upperCentralSeries e n, + ← e.toEquiv.forall_congr_right] + +namespace Group --- is_nilpotent is already defined in the root namespace (for elements of rings). +variable (G) in +-- `IsNilpotent` is already defined in the root namespace (for elements of rings). +-- TODO: Rename it to `IsNilpotentElement`? /-- A group `G` is nilpotent if its upper central series is eventually `G`. -/ -class Group.IsNilpotent (G : Type*) [Group G] : Prop where +@[mk_iff] +class IsNilpotent (G : Type*) [Group G] : Prop where nilpotent' : ∃ n : ℕ, upperCentralSeries G n = ⊤ -- Porting note: add lemma since infer kinds are unsupported in the definition of `IsNilpotent` -lemma Group.IsNilpotent.nilpotent (G : Type*) [Group G] [IsNilpotent G] : +lemma IsNilpotent.nilpotent (G : Type*) [Group G] [IsNilpotent G] : ∃ n : ℕ, upperCentralSeries G n = ⊤ := Group.IsNilpotent.nilpotent' -open Group +lemma isNilpotent_congr {H : Type*} [Group H] (e : G ≃* H) : IsNilpotent G ↔ IsNilpotent H := by + simp_rw [isNilpotent_iff] + refine exists_congr fun n ↦ ⟨fun h ↦ ?_, fun h ↦ ?_⟩ + · simp [← Subgroup.comap_top e.symm.toMonoidHom, ← h] + · simp [← Subgroup.comap_top e.toMonoidHom, ← h] -variable {G} +@[simp] lemma isNilpotent_top : IsNilpotent (⊤ : Subgroup G) ↔ IsNilpotent G := + isNilpotent_congr Subgroup.topEquiv + +variable (G) in +/-- A group `G` is virtually nilpotent if it has a nilpotent cofinite subgroup `N`. -/ +def IsVirtuallyNilpotent : Prop := ∃ N : Subgroup G, IsNilpotent N ∧ FiniteIndex N + +lemma IsNilpotent.isVirtuallyNilpotent (hG : IsNilpotent G) : IsVirtuallyNilpotent G := + ⟨⊤, by simpa, inferInstance⟩ + +end Group + +open Group /-- A sequence of subgroups of `G` is an ascending central series if `H 0` is trivial and `⁅H (n + 1), G⁆ ⊆ H n` for all `n`. Note that we do not require that `H n = G` for some `n`. -/ From cc5969c16900cd3fd7123e8c1045c8865b8ff1b0 Mon Sep 17 00:00:00 2001 From: Bhavik Mehta Date: Sat, 26 Oct 2024 16:21:20 +0000 Subject: [PATCH 458/505] feat(Matrix): characteristic polynomial of upper triangular matrix (#17338) Add simple lemmas relating block triangular matrices with characteristic polynomials. --- Mathlib/LinearAlgebra/Matrix/Block.lean | 12 ++++++++++ .../LinearAlgebra/Matrix/Charpoly/Basic.lean | 24 +++++++++++++++++++ 2 files changed, 36 insertions(+) diff --git a/Mathlib/LinearAlgebra/Matrix/Block.lean b/Mathlib/LinearAlgebra/Matrix/Block.lean index 8432d56b10b42..401caa2ef1eff 100644 --- a/Mathlib/LinearAlgebra/Matrix/Block.lean +++ b/Mathlib/LinearAlgebra/Matrix/Block.lean @@ -85,6 +85,18 @@ theorem BlockTriangular.add (hM : BlockTriangular M b) (hN : BlockTriangular N b theorem BlockTriangular.sub (hM : BlockTriangular M b) (hN : BlockTriangular N b) : BlockTriangular (M - N) b := fun i j h => by simp_rw [Matrix.sub_apply, hM h, hN h, sub_zero] +lemma BlockTriangular.add_iff_right (hM : BlockTriangular M b) : + BlockTriangular (M + N) b ↔ BlockTriangular N b := ⟨(by simpa using ·.sub hM), hM.add⟩ + +lemma BlockTriangular.add_iff_left (hN : BlockTriangular N b) : + BlockTriangular (M + N) b ↔ BlockTriangular M b := by rw [add_comm, hN.add_iff_right] + +lemma BlockTriangular.sub_iff_right (hM : BlockTriangular M b) : + BlockTriangular (M - N) b ↔ BlockTriangular N b := ⟨(by simpa using hM.sub ·), hM.sub⟩ + +lemma BlockTriangular.sub_iff_left (hN : BlockTriangular N b) : + BlockTriangular (M - N) b ↔ BlockTriangular M b := ⟨(by simpa using ·.add hN), (·.sub hN)⟩ + end LT section Preorder diff --git a/Mathlib/LinearAlgebra/Matrix/Charpoly/Basic.lean b/Mathlib/LinearAlgebra/Matrix/Charpoly/Basic.lean index 83b4e6ed3bec2..f2f8e2eff0c0c 100644 --- a/Mathlib/LinearAlgebra/Matrix/Charpoly/Basic.lean +++ b/Mathlib/LinearAlgebra/Matrix/Charpoly/Basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison -/ import Mathlib.LinearAlgebra.Matrix.Adjugate +import Mathlib.LinearAlgebra.Matrix.Block import Mathlib.RingTheory.PolynomialAlgebra /-! @@ -85,6 +86,16 @@ lemma charmatrix_fromBlocks : simp only [charmatrix] ext (i|i) (j|j) : 2 <;> simp [diagonal] +-- TODO: importing block triangular here is somewhat expensive, if more lemmas about it are added +-- to this file, it may be worth extracting things out to Charpoly/Block.lean +@[simp] +lemma charmatrix_blockTriangular_iff {α : Type*} [Preorder α] {M : Matrix n n R} {b : n → α} : + M.charmatrix.BlockTriangular b ↔ M.BlockTriangular b := by + rw [charmatrix, scalar_apply, RingHom.mapMatrix_apply, (blockTriangular_diagonal _).sub_iff_right] + simp [BlockTriangular] + +alias ⟨BlockTriangular.of_charmatrix, BlockTriangular.charmatrix⟩ := charmatrix_blockTriangular_iff + /-- The characteristic polynomial of a matrix `M` is given by $\det (t I - M)$. -/ def charpoly (M : Matrix n n R) : R[X] := @@ -112,6 +123,19 @@ lemma charpoly_fromBlocks_zero₂₁ : simp only [charpoly, charmatrix_fromBlocks, Matrix.map_zero _ (Polynomial.C_0), neg_zero, det_fromBlocks_zero₂₁] +lemma charmatrix_toSquareBlock {α : Type*} [DecidableEq α] {b : n → α} {a : α} : + (M.toSquareBlock b a).charmatrix = M.charmatrix.toSquareBlock b a := by + ext i j : 1 + simp [charmatrix_apply, toSquareBlock_def, diagonal_apply, Subtype.ext_iff] + +lemma BlockTriangular.charpoly {α : Type*} {b : n → α} [LinearOrder α] (h : M.BlockTriangular b) : + M.charpoly = ∏ a ∈ image b univ, (M.toSquareBlock b a).charpoly := by + simp only [Matrix.charpoly, h.charmatrix.det, charmatrix_toSquareBlock] + +lemma charpoly_of_upperTriangular [LinearOrder n] (M : Matrix n n R) (h : M.BlockTriangular id) : + M.charpoly = ∏ i : n, (X - C (M i i)) := by + simp [charpoly, det_of_upperTriangular h.charmatrix] + -- This proof follows http://drorbn.net/AcademicPensieve/2015-12/CayleyHamilton.pdf /-- The **Cayley-Hamilton Theorem**, that the characteristic polynomial of a matrix, applied to the matrix itself, is zero. From 2834fd211f1f6bccd9b3a792346ab1661de9ecb3 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Sat, 26 Oct 2024 17:00:22 +0000 Subject: [PATCH 459/505] refactor(RingTheory/Unramified): Switch official definition to allow general universes (#15135) Under this new definition, `Algebra.FormallySmooth` is only defined for commutative rings. Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- Mathlib.lean | 1 - Mathlib/RingTheory/Etale/Basic.lean | 30 ++-- Mathlib/RingTheory/Etale/Field.lean | 4 +- Mathlib/RingTheory/Unramified/Basic.lean | 144 ++++++++++++------ .../RingTheory/Unramified/Derivations.lean | 48 ------ Mathlib/RingTheory/Unramified/Field.lean | 4 +- Mathlib/RingTheory/Unramified/Finite.lean | 4 +- Mathlib/RingTheory/Unramified/Pi.lean | 2 +- 8 files changed, 117 insertions(+), 120 deletions(-) delete mode 100644 Mathlib/RingTheory/Unramified/Derivations.lean diff --git a/Mathlib.lean b/Mathlib.lean index 05c28e27f3830..ccfb1ea4dfdce 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4228,7 +4228,6 @@ import Mathlib.RingTheory.TwoSidedIdeal.Lattice import Mathlib.RingTheory.TwoSidedIdeal.Operations import Mathlib.RingTheory.UniqueFactorizationDomain import Mathlib.RingTheory.Unramified.Basic -import Mathlib.RingTheory.Unramified.Derivations import Mathlib.RingTheory.Unramified.Field import Mathlib.RingTheory.Unramified.Finite import Mathlib.RingTheory.Unramified.Pi diff --git a/Mathlib/RingTheory/Etale/Basic.lean b/Mathlib/RingTheory/Etale/Basic.lean index 554788c7bc998..155276c4c53a6 100644 --- a/Mathlib/RingTheory/Etale/Basic.lean +++ b/Mathlib/RingTheory/Etale/Basic.lean @@ -34,8 +34,8 @@ namespace Algebra section -variable (R : Type u) [CommSemiring R] -variable (A : Type u) [Semiring A] [Algebra R A] +variable (R : Type u) [CommRing R] +variable (A : Type u) [CommRing A] [Algebra R A] /-- An `R` algebra `A` is formally étale if for every `R`-algebra, every square-zero ideal `I : Ideal B` and `f : A →ₐ[R] B ⧸ I`, there exists exactly one lift `A →ₐ[R] B`. @@ -54,12 +54,12 @@ namespace FormallyEtale section -variable {R : Type u} [CommSemiring R] -variable {A : Type u} [Semiring A] [Algebra R A] +variable {R : Type u} [CommRing R] +variable {A : Type u} [CommRing A] [Algebra R A] theorem iff_unramified_and_smooth : FormallyEtale R A ↔ FormallyUnramified R A ∧ FormallySmooth R A := by - rw [formallyUnramified_iff, formallySmooth_iff, formallyEtale_iff] + rw [FormallyUnramified.iff_comp_injective, formallySmooth_iff, formallyEtale_iff] simp_rw [← forall_and, Function.Bijective] instance (priority := 100) to_unramified [h : FormallyEtale R A] : @@ -77,8 +77,8 @@ end section OfEquiv -variable {R : Type u} [CommSemiring R] -variable {A B : Type u} [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] +variable {R : Type u} [CommRing R] +variable {A B : Type u} [CommRing A] [Algebra R A] [CommRing B] [Algebra R B] theorem of_equiv [FormallyEtale R A] (e : A ≃ₐ[R] B) : FormallyEtale R B := FormallyEtale.iff_unramified_and_smooth.mpr @@ -91,9 +91,9 @@ end OfEquiv section Comp -variable (R : Type u) [CommSemiring R] -variable (A : Type u) [CommSemiring A] [Algebra R A] -variable (B : Type u) [Semiring B] [Algebra R B] [Algebra A B] [IsScalarTower R A B] +variable (R : Type u) [CommRing R] +variable (A : Type u) [CommRing A] [Algebra R A] +variable (B : Type u) [CommRing B] [Algebra R B] [Algebra A B] [IsScalarTower R A B] theorem comp [FormallyEtale R A] [FormallyEtale A B] : FormallyEtale R B := FormallyEtale.iff_unramified_and_smooth.mpr @@ -105,9 +105,9 @@ section BaseChange open scoped TensorProduct -variable {R : Type u} [CommSemiring R] -variable {A : Type u} [Semiring A] [Algebra R A] -variable (B : Type u) [CommSemiring B] [Algebra R B] +variable {R : Type u} [CommRing R] +variable {A : Type u} [CommRing A] [Algebra R A] +variable (B : Type u) [CommRing B] [Algebra R B] instance base_change [FormallyEtale R A] : FormallyEtale B (B ⊗[R] A) := FormallyEtale.iff_unramified_and_smooth.mpr ⟨inferInstance, inferInstance⟩ @@ -166,8 +166,8 @@ end FormallyEtale section -variable (R : Type u) [CommSemiring R] -variable (A : Type u) [Semiring A] [Algebra R A] +variable (R : Type u) [CommRing R] +variable (A : Type u) [CommRing A] [Algebra R A] /-- An `R`-algebra `A` is étale if it is formally étale and of finite presentation. diff --git a/Mathlib/RingTheory/Etale/Field.lean b/Mathlib/RingTheory/Etale/Field.lean index 8305070f2015e..937b07e0b9d8a 100644 --- a/Mathlib/RingTheory/Etale/Field.lean +++ b/Mathlib/RingTheory/Etale/Field.lean @@ -50,7 +50,7 @@ theorem of_isSeparable_aux [Algebra.IsSeparable K L] [EssFiniteType K L] : constructor -- We shall show that any `f : L → B/I` can be lifted to `L → B` if `I^2 = ⊥` intros B _ _ I h - refine ⟨(FormallyUnramified.of_isSeparable K L).1 I h, ?_⟩ + refine ⟨FormallyUnramified.iff_comp_injective.mp (FormallyUnramified.of_isSeparable K L) I h, ?_⟩ intro f -- By separability and finiteness, we may assume `L = K(α)` with `p` the minpoly of `α`. let pb := Field.powerBasisOfFiniteOfSeparable K L @@ -89,7 +89,7 @@ lemma of_isSeparable [Algebra.IsSeparable K L] : FormallyEtale K L := by -- We shall show that any `f : L → B/I` can be lifted to `L → B` if `I^2 = ⊥`. -- But we already know that there exists a unique lift for every finite subfield of `L` -- by `of_isSeparable_aux`, so we can glue them all together. - refine ⟨(FormallyUnramified.of_isSeparable K L).1 I h, ?_⟩ + refine ⟨FormallyUnramified.iff_comp_injective.mp (FormallyUnramified.of_isSeparable K L) I h, ?_⟩ intro f have : ∀ k : L, ∃! g : K⟮k⟯ →ₐ[K] B, (Ideal.Quotient.mkₐ K I).comp g = f.comp (IsScalarTower.toAlgHom K _ L) := by diff --git a/Mathlib/RingTheory/Unramified/Basic.lean b/Mathlib/RingTheory/Unramified/Basic.lean index cd3ce7546dad0..6b6f3781f7e4b 100644 --- a/Mathlib/RingTheory/Unramified/Basic.lean +++ b/Mathlib/RingTheory/Unramified/Basic.lean @@ -3,20 +3,19 @@ Copyright (c) 2022 Andrew Yang. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang -/ -import Mathlib.RingTheory.FinitePresentation import Mathlib.RingTheory.FiniteStability -import Mathlib.RingTheory.Localization.Away.Basic import Mathlib.RingTheory.Localization.Away.AdjoinRoot import Mathlib.RingTheory.QuotientNilpotent -import Mathlib.RingTheory.TensorProduct.Basic +import Mathlib.RingTheory.Kaehler.Basic /-! # Unramified morphisms -An `R`-algebra `A` is formally unramified if for every `R`-algebra, +An `R`-algebra `A` is formally unramified if `Ω[A⁄R]` is trivial. +This is equivalent to the standard definition "for every `R`-algebra, every square-zero ideal `I : Ideal B` and `f : A →ₐ[R] B ⧸ I`, there exists -at most one lift `A →ₐ[R] B`. +at most one lift `A →ₐ[R] B`". It is unramified if it is formally unramified and of finite type. Note that there are multiple definitions in the literature. The definition we give is equivalent to @@ -34,25 +33,28 @@ localization at an element. -- Porting note: added to make the syntax work below. open scoped TensorProduct -universe u +universe u v w namespace Algebra section -variable (R : Type u) [CommSemiring R] -variable (A : Type u) [Semiring A] [Algebra R A] +variable (R : Type v) [CommRing R] +variable (A : Type u) [CommRing A] [Algebra R A] -/-- An `R`-algebra `A` is formally unramified if for every `R`-algebra, every square-zero ideal -`I : Ideal B` and `f : A →ₐ[R] B ⧸ I`, there exists at most one lift `A →ₐ[R] B`. +/-- +An `R`-algebra `A` is formally unramified if `Ω[A⁄R]` is trivial. + +This is equivalent to "for every `R`-algebra, every square-zero ideal +`I : Ideal B` and `f : A →ₐ[R] B ⧸ I`, there exists at most one lift `A →ₐ[R] B`". +See `Algebra.FormallyUnramified.iff_comp_injective`. See . -/ @[mk_iff] class FormallyUnramified : Prop where - comp_injective : - ∀ ⦃B : Type u⦄ [CommRing B], - ∀ [Algebra R B] (I : Ideal B) (_ : I ^ 2 = ⊥), - Function.Injective ((Ideal.Quotient.mkₐ R I).comp : (A →ₐ[R] B) → A →ₐ[R] B ⧸ I) + subsingleton_kaehlerDifferential : Subsingleton (Ω[A⁄R]) + +attribute [instance] FormallyUnramified.subsingleton_kaehlerDifferential end @@ -60,16 +62,48 @@ namespace FormallyUnramified section -variable {R : Type u} [CommSemiring R] -variable {A : Type u} [Semiring A] [Algebra R A] -variable {B : Type u} [CommRing B] [Algebra R B] (I : Ideal B) - -theorem lift_unique {B : Type u} [CommRing B] [_RB : Algebra R B] +variable {R : Type v} [CommRing R] +variable {A : Type u} [CommRing A] [Algebra R A] +variable {B : Type w} [CommRing B] [Algebra R B] (I : Ideal B) + +theorem comp_injective [FormallyUnramified R A] (hI : I ^ 2 = ⊥) : + Function.Injective ((Ideal.Quotient.mkₐ R I).comp : (A →ₐ[R] B) → A →ₐ[R] B ⧸ I) := by + intro f₁ f₂ e + letI := f₁.toRingHom.toAlgebra + haveI := IsScalarTower.of_algebraMap_eq' f₁.comp_algebraMap.symm + have := + ((KaehlerDifferential.linearMapEquivDerivation R A).toEquiv.trans + (derivationToSquareZeroEquivLift I hI)).surjective.subsingleton + exact Subtype.ext_iff.mp (@Subsingleton.elim _ this ⟨f₁, rfl⟩ ⟨f₂, e.symm⟩) + +theorem iff_comp_injective : + FormallyUnramified R A ↔ + ∀ ⦃B : Type u⦄ [CommRing B], + ∀ [Algebra R B] (I : Ideal B) (_ : I ^ 2 = ⊥), + Function.Injective ((Ideal.Quotient.mkₐ R I).comp : (A →ₐ[R] B) → A →ₐ[R] B ⧸ I) := by + constructor + · intros; exact comp_injective _ ‹_› + · intro H + constructor + rw [← not_nontrivial_iff_subsingleton] + intro h + obtain ⟨f₁, f₂, e⟩ := (KaehlerDifferential.endEquiv R A).injective.nontrivial + apply e + ext1 + refine H + (RingHom.ker (TensorProduct.lmul' R (S := A)).kerSquareLift.toRingHom) ?_ ?_ + · rw [AlgHom.ker_kerSquareLift] + exact Ideal.cotangentIdeal_square _ + · ext x + apply RingHom.kerLift_injective (TensorProduct.lmul' R (S := A)).kerSquareLift.toRingHom + simpa using DFunLike.congr_fun (f₁.2.trans f₂.2.symm) x + +theorem lift_unique [FormallyUnramified R A] (I : Ideal B) (hI : IsNilpotent I) (g₁ g₂ : A →ₐ[R] B) (h : (Ideal.Quotient.mkₐ R I).comp g₁ = (Ideal.Quotient.mkₐ R I).comp g₂) : g₁ = g₂ := by revert g₁ g₂ change Function.Injective (Ideal.Quotient.mkₐ R I).comp - revert _RB + revert ‹Algebra R B› apply Ideal.IsNilpotent.induction_on (S := B) I hI · intro B _ I hI _; exact FormallyUnramified.comp_injective I hI · intro B _ I J hIJ h₁ h₂ _ g₁ g₂ e @@ -84,7 +118,7 @@ theorem ext [FormallyUnramified R A] (hI : IsNilpotent I) {g₁ g₂ : A →ₐ[ (H : ∀ x, Ideal.Quotient.mk I (g₁ x) = Ideal.Quotient.mk I (g₂ x)) : g₁ = g₂ := FormallyUnramified.lift_unique I hI g₁ g₂ (AlgHom.ext H) -theorem lift_unique_of_ringHom [FormallyUnramified R A] {C : Type u} [CommRing C] +theorem lift_unique_of_ringHom [FormallyUnramified R A] {C : Type*} [CommRing C] (f : B →+* C) (hf : IsNilpotent <| RingHom.ker f) (g₁ g₂ : A →ₐ[R] B) (h : f.comp ↑g₁ = f.comp (g₂ : A →+* B)) : g₁ = g₂ := FormallyUnramified.lift_unique _ hf _ _ @@ -94,32 +128,48 @@ theorem lift_unique_of_ringHom [FormallyUnramified R A] {C : Type u} [CommRing C simpa only [Ideal.Quotient.eq, Function.comp_apply, AlgHom.coe_comp, Ideal.Quotient.mkₐ_eq_mk, RingHom.mem_ker, map_sub, sub_eq_zero]) -theorem ext' [FormallyUnramified R A] {C : Type u} [CommRing C] (f : B →+* C) +theorem ext' [FormallyUnramified R A] {C : Type*} [CommRing C] (f : B →+* C) (hf : IsNilpotent <| RingHom.ker f) (g₁ g₂ : A →ₐ[R] B) (h : ∀ x, f (g₁ x) = f (g₂ x)) : g₁ = g₂ := FormallyUnramified.lift_unique_of_ringHom f hf g₁ g₂ (RingHom.ext h) -theorem lift_unique' [FormallyUnramified R A] {C : Type u} [CommRing C] +theorem lift_unique' [FormallyUnramified R A] {C : Type*} [CommRing C] [Algebra R C] (f : B →ₐ[R] C) (hf : IsNilpotent <| RingHom.ker (f : B →+* C)) (g₁ g₂ : A →ₐ[R] B) (h : f.comp g₁ = f.comp g₂) : g₁ = g₂ := FormallyUnramified.ext' _ hf g₁ g₂ (AlgHom.congr_fun h) -instance : FormallyUnramified R R := by - constructor +end + +instance {R : Type*} [CommRing R] : FormallyUnramified R R := by + rw [iff_comp_injective] intros B _ _ _ _ f₁ f₂ _ exact Subsingleton.elim _ _ -end +section OfEquiv + +variable {R : Type*} [CommRing R] +variable {A B : Type*} [CommRing A] [Algebra R A] [CommRing B] [Algebra R B] + +theorem of_equiv [FormallyUnramified R A] (e : A ≃ₐ[R] B) : + FormallyUnramified R B := by + rw [iff_comp_injective] + intro C _ _ I hI f₁ f₂ e' + rw [← f₁.comp_id, ← f₂.comp_id, ← e.comp_symm, ← AlgHom.comp_assoc, ← AlgHom.comp_assoc] + congr 1 + refine FormallyUnramified.comp_injective I hI ?_ + rw [← AlgHom.comp_assoc, e', AlgHom.comp_assoc] + +end OfEquiv section Comp -variable (R : Type u) [CommSemiring R] -variable (A : Type u) [CommSemiring A] [Algebra R A] -variable (B : Type u) [Semiring B] [Algebra R B] [Algebra A B] [IsScalarTower R A B] +variable (R : Type*) [CommRing R] +variable (A : Type*) [CommRing A] [Algebra R A] +variable (B : Type*) [CommRing B] [Algebra R B] [Algebra A B] [IsScalarTower R A B] theorem comp [FormallyUnramified R A] [FormallyUnramified A B] : FormallyUnramified R B := by - constructor + rw [iff_comp_injective] intro C _ _ I hI f₁ f₂ e have e' := FormallyUnramified.lift_unique I ⟨2, hI⟩ (f₁.comp <| IsScalarTower.toAlgHom R A B) @@ -133,7 +183,7 @@ theorem comp [FormallyUnramified R A] [FormallyUnramified A B] : exact FormallyUnramified.ext I ⟨2, hI⟩ (AlgHom.congr_fun e) theorem of_comp [FormallyUnramified R B] : FormallyUnramified A B := by - constructor + rw [iff_comp_injective] intro Q _ _ I e f₁ f₂ e' letI := ((algebraMap A Q).comp (algebraMap R A)).toAlgebra letI : IsScalarTower R A Q := IsScalarTower.of_algebraMap_eq' rfl @@ -146,13 +196,13 @@ end Comp section of_surjective -variable {R : Type u} [CommSemiring R] -variable {A B : Type u} [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] +variable {R : Type*} [CommRing R] +variable {A B : Type*} [CommRing A] [Algebra R A] [CommRing B] [Algebra R B] /-- This holds in general for epimorphisms. -/ theorem of_surjective [FormallyUnramified R A] (f : A →ₐ[R] B) (H : Function.Surjective f) : FormallyUnramified R B := by - constructor + rw [iff_comp_injective] intro Q _ _ I hI f₁ f₂ e ext x obtain ⟨x, rfl⟩ := H x @@ -163,11 +213,7 @@ theorem of_surjective [FormallyUnramified R A] (f : A →ₐ[R] B) (H : Function instance quotient {A} [CommRing A] [Algebra R A] [FormallyUnramified R A] (I : Ideal A) : FormallyUnramified R (A ⧸ I) := - FormallyUnramified.of_surjective (IsScalarTower.toAlgHom _ _ _) Ideal.Quotient.mk_surjective - -theorem of_equiv [FormallyUnramified R A] (e : A ≃ₐ[R] B) : - FormallyUnramified R B := - of_surjective e.toAlgHom e.surjective + FormallyUnramified.of_surjective (IsScalarTower.toAlgHom R A (A ⧸ I)) Ideal.Quotient.mk_surjective theorem iff_of_equiv (e : A ≃ₐ[R] B) : FormallyUnramified R A ↔ FormallyUnramified R B := ⟨fun _ ↦ of_equiv e, fun _ ↦ of_equiv e.symm⟩ @@ -178,13 +224,13 @@ section BaseChange open scoped TensorProduct -variable {R : Type u} [CommSemiring R] -variable {A : Type u} [Semiring A] [Algebra R A] -variable (B : Type u) [CommSemiring B] [Algebra R B] +variable {R : Type*} [CommRing R] +variable {A : Type*} [CommRing A] [Algebra R A] +variable (B : Type*) [CommRing B] [Algebra R B] instance base_change [FormallyUnramified R A] : FormallyUnramified B (B ⊗[R] A) := by - constructor + rw [iff_comp_injective] intro C _ _ I hI f₁ f₂ e letI := ((algebraMap B C).comp (algebraMap R B)).toAlgebra haveI : IsScalarTower R B C := IsScalarTower.of_algebraMap_eq' rfl @@ -196,7 +242,7 @@ end BaseChange section Localization -variable {R S Rₘ Sₘ : Type u} [CommRing R] [CommRing S] [CommRing Rₘ] [CommRing Sₘ] +variable {R S Rₘ Sₘ : Type*} [CommRing R] [CommRing S] [CommRing Rₘ] [CommRing Sₘ] variable (M : Submonoid R) variable [Algebra R S] [Algebra R Sₘ] [Algebra S Sₘ] [Algebra R Rₘ] [Algebra Rₘ Sₘ] variable [IsScalarTower R Rₘ Sₘ] [IsScalarTower R S Sₘ] @@ -208,7 +254,7 @@ include M /-- This holds in general for epimorphisms. -/ theorem of_isLocalization [IsLocalization M Rₘ] : FormallyUnramified R Rₘ := by - constructor + rw [iff_comp_injective] intro Q _ _ I _ f₁ f₂ _ apply AlgHom.coe_ringHom_injective refine IsLocalization.ringHom_ext M ?_ @@ -239,8 +285,8 @@ end FormallyUnramified section -variable (R : Type u) [CommSemiring R] -variable (A : Type u) [Semiring A] [Algebra R A] +variable (R : Type*) [CommRing R] +variable (A : Type*) [CommRing A] [Algebra R A] /-- An `R`-algebra `A` is unramified if it is formally unramified and of finite type. @@ -258,8 +304,8 @@ namespace Unramified attribute [instance] formallyUnramified finiteType -variable {R : Type u} [CommRing R] -variable {A B : Type u} [CommRing A] [Algebra R A] [CommRing B] [Algebra R B] +variable {R : Type*} [CommRing R] +variable {A B : Type*} [CommRing A] [Algebra R A] [CommRing B] [Algebra R B] /-- Being unramified is transported via algebra isomorphisms. -/ theorem of_equiv [Unramified R A] (e : A ≃ₐ[R] B) : Unramified R B where diff --git a/Mathlib/RingTheory/Unramified/Derivations.lean b/Mathlib/RingTheory/Unramified/Derivations.lean deleted file mode 100644 index 76feabb3ef6be..0000000000000 --- a/Mathlib/RingTheory/Unramified/Derivations.lean +++ /dev/null @@ -1,48 +0,0 @@ -/- -Copyright (c) 2022 Andrew Yang. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Andrew Yang --/ -import Mathlib.RingTheory.Kaehler.Basic -import Mathlib.RingTheory.Unramified.Basic - -/-! - -# Differential properties of formally unramified algebras - -We show that `R`-algebra `A` is formally unramified iff the Kaehler differentials vanish. - --/ - -universe u - -namespace Algebra - -variable {R S : Type u} [CommRing R] [CommRing S] [Algebra R S] - -instance FormallyUnramified.subsingleton_kaehlerDifferential [FormallyUnramified R S] : - Subsingleton (Ω[S⁄R]) := by - rw [← not_nontrivial_iff_subsingleton] - intro h - obtain ⟨f₁, f₂, e⟩ := (KaehlerDifferential.endEquiv R S).injective.nontrivial - apply e - ext1 - apply FormallyUnramified.lift_unique' _ _ _ _ (f₁.2.trans f₂.2.symm) - rw [← AlgHom.toRingHom_eq_coe, AlgHom.ker_kerSquareLift] - exact ⟨_, Ideal.cotangentIdeal_square _⟩ - -theorem FormallyUnramified.iff_subsingleton_kaehlerDifferential : - FormallyUnramified R S ↔ Subsingleton (Ω[S⁄R]) := by - constructor - · intros; infer_instance - · intro H - constructor - intro B _ _ I hI f₁ f₂ e - letI := f₁.toRingHom.toAlgebra - haveI := IsScalarTower.of_algebraMap_eq' f₁.comp_algebraMap.symm - have := - ((KaehlerDifferential.linearMapEquivDerivation R S).toEquiv.trans - (derivationToSquareZeroEquivLift I hI)).surjective.subsingleton - exact Subtype.ext_iff.mp (@Subsingleton.elim _ this ⟨f₁, rfl⟩ ⟨f₂, e.symm⟩) - -end Algebra diff --git a/Mathlib/RingTheory/Unramified/Field.lean b/Mathlib/RingTheory/Unramified/Field.lean index 3a0152d738c1c..b2faae3e80dec 100644 --- a/Mathlib/RingTheory/Unramified/Field.lean +++ b/Mathlib/RingTheory/Unramified/Field.lean @@ -40,7 +40,7 @@ open scoped TensorProduct namespace Algebra.FormallyUnramified theorem of_isSeparable [Algebra.IsSeparable K L] : FormallyUnramified K L := by - constructor + rw [iff_comp_injective] intros B _ _ I hI f₁ f₂ e ext x have : f₁ x - f₂ x ∈ I := by @@ -206,7 +206,7 @@ theorem isSeparable : Algebra.IsSeparable K L := by rw [← range_eq_top_of_isPurelyInseparable (separableClosure K L) L] simp -theorem iff_isSeparable (L) [Field L] [Algebra K L] [EssFiniteType K L] : +theorem iff_isSeparable (L : Type u) [Field L] [Algebra K L] [EssFiniteType K L] : FormallyUnramified K L ↔ Algebra.IsSeparable K L := ⟨fun _ ↦ isSeparable K L, fun _ ↦ of_isSeparable K L⟩ diff --git a/Mathlib/RingTheory/Unramified/Finite.lean b/Mathlib/RingTheory/Unramified/Finite.lean index 8652ea0e4337f..59c67331b55c7 100644 --- a/Mathlib/RingTheory/Unramified/Finite.lean +++ b/Mathlib/RingTheory/Unramified/Finite.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang -/ import Mathlib.RingTheory.Ideal.IdempotentFG -import Mathlib.RingTheory.Unramified.Derivations +import Mathlib.RingTheory.Unramified.Basic import Mathlib.RingTheory.Flat.Stability /-! @@ -48,7 +48,7 @@ A finite-type `R`-algebra `S` is (formally) unramified iff there exists a `t : S theorem iff_exists_tensorProduct [EssFiniteType R S] : FormallyUnramified R S ↔ ∃ t : S ⊗[R] S, (∀ s, ((1 : S) ⊗ₜ[R] s - s ⊗ₜ[R] (1 : S)) * t = 0) ∧ TensorProduct.lmul' R t = 1 := by - rw [iff_subsingleton_kaehlerDifferential, KaehlerDifferential, + rw [formallyUnramified_iff, KaehlerDifferential, Ideal.cotangent_subsingleton_iff, Ideal.isIdempotentElem_iff_of_fg _ (KaehlerDifferential.ideal_fg R S)] have : ∀ t : S ⊗[R] S, TensorProduct.lmul' R t = 1 ↔ 1 - t ∈ KaehlerDifferential.ideal R S := by diff --git a/Mathlib/RingTheory/Unramified/Pi.lean b/Mathlib/RingTheory/Unramified/Pi.lean index dc1d2ec8c816b..755c16a601f58 100644 --- a/Mathlib/RingTheory/Unramified/Pi.lean +++ b/Mathlib/RingTheory/Unramified/Pi.lean @@ -31,7 +31,7 @@ theorem pi_iff : · intro _ i exact FormallyUnramified.of_surjective (Pi.evalAlgHom R f i) (Function.surjective_eval i) · intro H - constructor + rw [iff_comp_injective] intros B _ _ J hJ f₁ f₂ e ext g rw [← Finset.univ_sum_single g, map_sum, map_sum] From 8102abc0537f325d1a06a2a107b481799cfd2877 Mon Sep 17 00:00:00 2001 From: Seewoo Lee Date: Sat, 26 Oct 2024 17:47:25 +0000 Subject: [PATCH 460/505] feat(RingTheory/Radical): Theorems for coprime elements + alpha (#15331) Co-authored-by: Jineon Baek Co-authored-by: Seewoo Lee <49933279+seewoo5@users.noreply.github.com> --- Mathlib/RingTheory/Radical.lean | 59 +++++++++++++++++++++++++++++++++ 1 file changed, 59 insertions(+) diff --git a/Mathlib/RingTheory/Radical.lean b/Mathlib/RingTheory/Radical.lean index 0c28c161cdf4f..0e8777e07f247 100644 --- a/Mathlib/RingTheory/Radical.lean +++ b/Mathlib/RingTheory/Radical.lean @@ -108,4 +108,63 @@ theorem radical_pow_of_prime {a : M} (ha : Prime a) {n : ℕ} (hn : 0 < n) : rw [radical_pow a hn] exact radical_of_prime ha +theorem radical_ne_zero (a : M) [Nontrivial M] : radical a ≠ 0 := by + rw [radical, ← Finset.prod_val] + apply Multiset.prod_ne_zero + rw [primeFactors] + simp only [Multiset.toFinset_val, Multiset.mem_dedup] + exact zero_not_mem_normalizedFactors _ + end UniqueFactorizationMonoid + +open UniqueFactorizationMonoid + +namespace UniqueFactorizationDomain +-- Theorems for UFDs + +variable {R : Type*} [CommRing R] [IsDomain R] [NormalizationMonoid R] + [UniqueFactorizationMonoid R] + +/-- Coprime elements have disjoint prime factors (as multisets). -/ +theorem disjoint_normalizedFactors {a b : R} (hc : IsCoprime a b) : + (normalizedFactors a).Disjoint (normalizedFactors b) := by + intro x hxa hxb + have x_dvd_a := dvd_of_mem_normalizedFactors hxa + have x_dvd_b := dvd_of_mem_normalizedFactors hxb + have xp := prime_of_normalized_factor x hxa + exact xp.not_unit (hc.isUnit_of_dvd' x_dvd_a x_dvd_b) + +/-- Coprime elements have disjoint prime factors (as finsets). -/ +theorem disjoint_primeFactors {a b : R} (hc : IsCoprime a b) : + Disjoint (primeFactors a) (primeFactors b) := + Multiset.disjoint_toFinset.mpr (disjoint_normalizedFactors hc) + +theorem mul_primeFactors_disjUnion {a b : R} (ha : a ≠ 0) (hb : b ≠ 0) + (hc : IsCoprime a b) : + primeFactors (a * b) = + (primeFactors a).disjUnion (primeFactors b) (disjoint_primeFactors hc) := by + rw [Finset.disjUnion_eq_union] + simp_rw [primeFactors] + rw [normalizedFactors_mul ha hb, Multiset.toFinset_add] + +@[simp] +theorem radical_neg_one : radical (-1 : R) = 1 := + radical_unit_eq_one isUnit_one.neg + +/-- Radical is multiplicative for coprime elements. -/ +theorem radical_mul {a b : R} (hc : IsCoprime a b) : + radical (a * b) = (radical a) * (radical b) := by + by_cases ha : a = 0 + · subst ha; rw [isCoprime_zero_left] at hc + simp only [zero_mul, radical_zero_eq, one_mul, radical_unit_eq_one hc] + by_cases hb : b = 0 + · subst hb; rw [isCoprime_zero_right] at hc + simp only [mul_zero, radical_zero_eq, mul_one, radical_unit_eq_one hc] + simp_rw [radical] + rw [mul_primeFactors_disjUnion ha hb hc] + rw [Finset.prod_disjUnion (disjoint_primeFactors hc)] + +theorem radical_neg {a : R} : radical (-a) = radical a := + radical_eq_of_associated Associated.rfl.neg_left + +end UniqueFactorizationDomain From 991921d3bd6203a88526aed794553b0ba2350735 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Sat, 26 Oct 2024 18:11:53 +0000 Subject: [PATCH 461/505] feat: define `cfcHomSuperset` (#18126) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This defines, for convenience, a map `cfcHomSuperset : C(s, R) → A` whenever `spectrum R a ⊆ s` in the natural way (i.e., by precomposition with the inclusion map). --- .../NonUnital.lean | 44 +++++++++++++++++++ .../ContinuousFunctionalCalculus/Unital.lean | 29 ++++++++++++ 2 files changed, 73 insertions(+) diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean index 458c86fac7e39..3d40d1c59a022 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean @@ -658,6 +658,50 @@ end Ring end Order +/-! ### `cfcₙHom` on a superset of the quasispectrum -/ + +section Superset + +open ContinuousMapZero + +variable {R A : Type*} {p : A → Prop} [CommSemiring R] [Nontrivial R] [StarRing R] + [MetricSpace R] [TopologicalSemiring R] [ContinuousStar R] [NonUnitalRing A] [StarRing A] + [TopologicalSpace A] [Module R A] [IsScalarTower R A A] [SMulCommClass R A A] + [instCFCₙ : NonUnitalContinuousFunctionalCalculus R p] + +/-- The composition of `cfcₙHom` with the natural embedding `C(s, R)₀ → C(quasispectrum R a, R)₀` +whenever `quasispectrum R a ⊆ s`. + +This is sometimes necessary in order to consider the same continuous functions applied to multiple +distinct elements, with the added constraint that `cfcₙ` does not suffice. This can occur, for +example, if it is necessary to use uniqueness of this continuous functional calculus. A practical +example can be found in the proof of `CFC.posPart_negPart_unique`. -/ +@[simps!] +noncomputable def cfcₙHomSuperset {a : A} (ha : p a) {s : Set R} (hs : σₙ R a ⊆ s) : + letI : Zero s := ⟨0, hs (quasispectrum.zero_mem R a)⟩ + C(s, R)₀ →⋆ₙₐ[R] A := + letI : Zero s := ⟨0, hs (quasispectrum.zero_mem R a)⟩ + cfcₙHom ha (R := R) |>.comp <| ContinuousMapZero.nonUnitalStarAlgHom_precomp R <| + ⟨⟨_, continuous_id.subtype_map hs⟩, rfl⟩ + +lemma cfcₙHomSuperset_continuous {a : A} (ha : p a) {s : Set R} (hs : σₙ R a ⊆ s) : + Continuous (cfcₙHomSuperset ha hs) := + letI : Zero s := ⟨0, hs (quasispectrum.zero_mem R a)⟩ + (cfcₙHom_continuous ha).comp <| ContinuousMapZero.continuous_comp_left _ + +lemma cfcₙHomSuperset_id {a : A} (ha : p a) {s : Set R} (hs : σₙ R a ⊆ s) : + letI : Zero s := ⟨0, hs (quasispectrum.zero_mem R a)⟩ + cfcₙHomSuperset ha hs ⟨.restrict s <| .id R, rfl⟩ = a := + cfcₙHom_id ha + +/-- this version uses `ContinuousMapZero.id`. -/ +lemma cfcₙHomSuperset_id' {a : A} (ha : p a) {s : Set R} (hs : σₙ R a ⊆ s) : + letI : Zero s := ⟨0, hs (quasispectrum.zero_mem R a)⟩ + cfcₙHomSuperset ha hs (.id rfl) = a := + cfcₙHom_id ha + +end Superset + /-! ### Obtain a non-unital continuous functional calculus from a unital one -/ section UnitalToNonUnital diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean index 820b45e9bdb3c..6e8de01f4cc2d 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean @@ -1034,3 +1034,32 @@ lemma CFC.one_le_iff (a : A) (ha : p a := by cfc_tac) : end Ring end Order + +/-! ### `cfcHom` on a superset of the spectrum -/ + +section Superset + +variable {R A : Type*} {p : A → Prop} [CommSemiring R] [StarRing R] + [MetricSpace R] [TopologicalSemiring R] [ContinuousStar R] [Ring A] [StarRing A] + [TopologicalSpace A] [Algebra R A] [instCFC : ContinuousFunctionalCalculus R p] + +/-- The composition of `cfcHom` with the natural embedding `C(s, R) → C(spectrum R a, R)` +whenever `spectrum R a ⊆ s`. + +This is sometimes necessary in order to consider the same continuous functions applied to multiple +distinct elements, with the added constraint that `cfc` does not suffice. This can occur, for +example, if it is necessary to use uniqueness of this continuous functional calculus. -/ +@[simps!] +noncomputable def cfcHomSuperset {a : A} (ha : p a) {s : Set R} (hs : spectrum R a ⊆ s) : + C(s, R) →⋆ₐ[R] A := + cfcHom ha |>.comp <| ContinuousMap.compStarAlgHom' R R <| ⟨_, continuous_id.subtype_map hs⟩ + +lemma cfcHomSuperset_continuous {a : A} (ha : p a) {s : Set R} (hs : spectrum R a ⊆ s) : + Continuous (cfcHomSuperset ha hs) := + (cfcHom_continuous ha).comp <| ContinuousMap.continuous_precomp _ + +lemma cfcHomSuperset_id {a : A} (ha : p a) {s : Set R} (hs : spectrum R a ⊆ s) : + cfcHomSuperset ha hs (.restrict s <| .id R) = a := + cfcHom_id ha + +end Superset From 0fb334674969d63019f5ae82ebac5e2ff1cd4b67 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Calle=20S=C3=B6nne?= Date: Sat, 26 Oct 2024 22:53:57 +0000 Subject: [PATCH 462/505] chore(Bicategory): restructure oplax modifications (#18250) This PR moves Oplax modifications to its own file `Modification/Oplax.lean` and also moves `FunctorBicategory.lean` to `FunctorBicategory/Oplax.lean`. This will improve the file structure for later PRs generalizing these constructions to pseudofunctors. --- Mathlib.lean | 3 +- .../Oplax.lean} | 4 +- .../Bicategory/Modification/Oplax.lean | 134 ++++++++++++++++++ .../NaturalTransformation/Oplax.lean | 104 -------------- 4 files changed, 138 insertions(+), 107 deletions(-) rename Mathlib/CategoryTheory/Bicategory/{FunctorBicategory.lean => FunctorBicategory/Oplax.lean} (96%) create mode 100644 Mathlib/CategoryTheory/Bicategory/Modification/Oplax.lean diff --git a/Mathlib.lean b/Mathlib.lean index ccfb1ea4dfdce..01c450dd99fe5 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1495,12 +1495,13 @@ import Mathlib.CategoryTheory.Bicategory.Functor.LocallyDiscrete import Mathlib.CategoryTheory.Bicategory.Functor.Oplax import Mathlib.CategoryTheory.Bicategory.Functor.Prelax import Mathlib.CategoryTheory.Bicategory.Functor.Pseudofunctor -import Mathlib.CategoryTheory.Bicategory.FunctorBicategory +import Mathlib.CategoryTheory.Bicategory.FunctorBicategory.Oplax import Mathlib.CategoryTheory.Bicategory.Grothendieck import Mathlib.CategoryTheory.Bicategory.Kan.Adjunction import Mathlib.CategoryTheory.Bicategory.Kan.HasKan import Mathlib.CategoryTheory.Bicategory.Kan.IsKan import Mathlib.CategoryTheory.Bicategory.LocallyDiscrete +import Mathlib.CategoryTheory.Bicategory.Modification.Oplax import Mathlib.CategoryTheory.Bicategory.NaturalTransformation.Oplax import Mathlib.CategoryTheory.Bicategory.NaturalTransformation.Strong import Mathlib.CategoryTheory.Bicategory.SingleObj diff --git a/Mathlib/CategoryTheory/Bicategory/FunctorBicategory.lean b/Mathlib/CategoryTheory/Bicategory/FunctorBicategory/Oplax.lean similarity index 96% rename from Mathlib/CategoryTheory/Bicategory/FunctorBicategory.lean rename to Mathlib/CategoryTheory/Bicategory/FunctorBicategory/Oplax.lean index 9822588b0753f..02482d12bfd09 100644 --- a/Mathlib/CategoryTheory/Bicategory/FunctorBicategory.lean +++ b/Mathlib/CategoryTheory/Bicategory/FunctorBicategory/Oplax.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Yuma Mizuno. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yuma Mizuno -/ -import Mathlib.CategoryTheory.Bicategory.NaturalTransformation.Oplax +import Mathlib.CategoryTheory.Bicategory.Modification.Oplax /-! # The bicategory of oplax functors between two bicategories @@ -17,7 +17,7 @@ Given bicategories `B` and `C`, we give a bicategory structure on `OplaxFunctor namespace CategoryTheory -open Category Bicategory +open Category Bicategory Oplax open scoped Bicategory diff --git a/Mathlib/CategoryTheory/Bicategory/Modification/Oplax.lean b/Mathlib/CategoryTheory/Bicategory/Modification/Oplax.lean new file mode 100644 index 0000000000000..88fc6d73f8f0f --- /dev/null +++ b/Mathlib/CategoryTheory/Bicategory/Modification/Oplax.lean @@ -0,0 +1,134 @@ +/- +Copyright (c) 2024 Calle Sönne. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yuma Mizuno, Calle Sönne +-/ + +import Mathlib.CategoryTheory.Bicategory.NaturalTransformation.Oplax + +/-! +# Modifications between oplax transformations + +A modification `Γ` between oplax transformations `η` and `θ` consists of a family of +2-morphisms `Γ.app a : η.app a ⟶ θ.app a`, which satisfies the equation +`(F.map f ◁ app b) ≫ θ.naturality f = η.naturality f ≫ (app a ▷ G.map f)` +for each 1-morphism `f : a ⟶ b`. + +## Main definitions + +* `Modification η θ` : modifications between oplax transformations `η` and `θ` +* `Modification.vcomp η θ` : the vertical composition of oplax transformations `η` + and `θ` +* `OplaxTrans.category F G` : the category structure on the oplax transformations + between `F` and `G` + +-/ + +namespace CategoryTheory.Oplax + +open Category Bicategory + +universe w₁ w₂ v₁ v₂ u₁ u₂ + +variable {B : Type u₁} [Bicategory.{w₁, v₁} B] {C : Type u₂} [Bicategory.{w₂, v₂} C] + {F G : OplaxFunctor B C} (η θ : F ⟶ G) + +variable {F G : OplaxFunctor B C} + +/-- A modification `Γ` between oplax natural transformations `η` and `θ` consists of a family of +2-morphisms `Γ.app a : η.app a ⟶ θ.app a`, which satisfies the equation +`(F.map f ◁ app b) ≫ θ.naturality f = η.naturality f ≫ (app a ▷ G.map f)` +for each 1-morphism `f : a ⟶ b`. +-/ +@[ext] +structure Modification (η θ : F ⟶ G) where + /-- The underlying family of 2-morphism. -/ + app (a : B) : η.app a ⟶ θ.app a + /-- The naturality condition. -/ + naturality : + ∀ {a b : B} (f : a ⟶ b), + F.map f ◁ app b ≫ θ.naturality f = η.naturality f ≫ app a ▷ G.map f := by + aesop_cat + +attribute [reassoc (attr := simp)] Modification.naturality + +variable {η θ ι : F ⟶ G} + +namespace Modification + +variable (η) + +/-- The identity modification. -/ +@[simps] +def id : Modification η η where app a := 𝟙 (η.app a) + +instance : Inhabited (Modification η η) := + ⟨Modification.id η⟩ + +variable {η} + +section + +variable (Γ : Modification η θ) {a b c : B} {a' : C} + +@[reassoc (attr := simp)] +theorem whiskerLeft_naturality (f : a' ⟶ F.obj b) (g : b ⟶ c) : + f ◁ F.map g ◁ Γ.app c ≫ f ◁ θ.naturality g = f ◁ η.naturality g ≫ f ◁ Γ.app b ▷ G.map g := by + simp_rw [← Bicategory.whiskerLeft_comp, naturality] + +@[reassoc (attr := simp)] +theorem whiskerRight_naturality (f : a ⟶ b) (g : G.obj b ⟶ a') : + F.map f ◁ Γ.app b ▷ g ≫ (α_ _ _ _).inv ≫ θ.naturality f ▷ g = + (α_ _ _ _).inv ≫ η.naturality f ▷ g ≫ Γ.app a ▷ G.map f ▷ g := by + simp_rw [associator_inv_naturality_middle_assoc, ← comp_whiskerRight, naturality] + +end + +/-- Vertical composition of modifications. -/ +@[simps] +def vcomp (Γ : Modification η θ) (Δ : Modification θ ι) : Modification η ι where + app a := Γ.app a ≫ Δ.app a + +end Modification + +/-- Category structure on the oplax natural transformations between OplaxFunctors. -/ +@[simps] +instance category (F G : OplaxFunctor B C) : Category (F ⟶ G) where + Hom := Modification + id := Modification.id + comp := Modification.vcomp + +@[ext] +lemma ext {F G : OplaxFunctor B C} {α β : F ⟶ G} {m n : α ⟶ β} (w : ∀ b, m.app b = n.app b) : + m = n := by + apply Modification.ext + ext + apply w + +/-- Version of `Modification.id_app` using category notation -/ +@[simp] +lemma Modification.id_app' {X : B} {F G : OplaxFunctor B C} (α : F ⟶ G) : + Modification.app (𝟙 α) X = 𝟙 (α.app X) := rfl + +/-- Version of `Modification.comp_app` using category notation -/ +@[simp] +lemma Modification.comp_app' {X : B} {F G : OplaxFunctor B C} {α β γ : F ⟶ G} + (m : α ⟶ β) (n : β ⟶ γ) : (m ≫ n).app X = m.app X ≫ n.app X := + rfl + +/-- Construct a modification isomorphism between oplax natural transformations +by giving object level isomorphisms, and checking naturality only in the forward direction. +-/ +@[simps] +def ModificationIso.ofComponents (app : ∀ a, η.app a ≅ θ.app a) + (naturality : + ∀ {a b} (f : a ⟶ b), + F.map f ◁ (app b).hom ≫ θ.naturality f = η.naturality f ≫ (app a).hom ▷ G.map f) : + η ≅ θ where + hom := { app := fun a => (app a).hom } + inv := + { app := fun a => (app a).inv + naturality := fun {a b} f => by + simpa using congr_arg (fun f => _ ◁ (app b).inv ≫ f ≫ (app a).inv ▷ _) (naturality f).symm } + +end CategoryTheory.Oplax diff --git a/Mathlib/CategoryTheory/Bicategory/NaturalTransformation/Oplax.lean b/Mathlib/CategoryTheory/Bicategory/NaturalTransformation/Oplax.lean index ee6b3cb99b45e..0d8ff2f83325f 100644 --- a/Mathlib/CategoryTheory/Bicategory/NaturalTransformation/Oplax.lean +++ b/Mathlib/CategoryTheory/Bicategory/NaturalTransformation/Oplax.lean @@ -174,110 +174,6 @@ instance : CategoryStruct (OplaxFunctor B C) where end -section - -variable {F G : OplaxFunctor B C} - -/-- A modification `Γ` between oplax natural transformations `η` and `θ` consists of a family of -2-morphisms `Γ.app a : η.app a ⟶ θ.app a`, which satisfies the equation -`(F.map f ◁ app b) ≫ θ.naturality f = η.naturality f ≫ (app a ▷ G.map f)` -for each 1-morphism `f : a ⟶ b`. --/ -@[ext] -structure Modification (η θ : F ⟶ G) where - app (a : B) : η.app a ⟶ θ.app a - naturality : - ∀ {a b : B} (f : a ⟶ b), - F.map f ◁ app b ≫ θ.naturality f = η.naturality f ≫ app a ▷ G.map f := by - aesop_cat - -attribute [nolint docBlame] CategoryTheory.OplaxNatTrans.Modification.app - CategoryTheory.OplaxNatTrans.Modification.naturality - -/- Porting note: removed primes from field names and removed `restate_axiom` since that is no longer - needed in Lean 4 -/ - -attribute [reassoc (attr := simp)] Modification.naturality - -variable {η θ ι : F ⟶ G} - -namespace Modification - -variable (η) - -/-- The identity modification. -/ -@[simps] -def id : Modification η η where app a := 𝟙 (η.app a) - -instance : Inhabited (Modification η η) := - ⟨Modification.id η⟩ - -variable {η} - -section - -variable (Γ : Modification η θ) {a b c : B} {a' : C} - -@[reassoc (attr := simp)] -theorem whiskerLeft_naturality (f : a' ⟶ F.obj b) (g : b ⟶ c) : - f ◁ F.map g ◁ Γ.app c ≫ f ◁ θ.naturality g = f ◁ η.naturality g ≫ f ◁ Γ.app b ▷ G.map g := by - simp_rw [← Bicategory.whiskerLeft_comp, naturality] - -@[reassoc (attr := simp)] -theorem whiskerRight_naturality (f : a ⟶ b) (g : G.obj b ⟶ a') : - F.map f ◁ Γ.app b ▷ g ≫ (α_ _ _ _).inv ≫ θ.naturality f ▷ g = - (α_ _ _ _).inv ≫ η.naturality f ▷ g ≫ Γ.app a ▷ G.map f ▷ g := by - simp_rw [associator_inv_naturality_middle_assoc, ← comp_whiskerRight, naturality] - -end - -/-- Vertical composition of modifications. -/ -@[simps] -def vcomp (Γ : Modification η θ) (Δ : Modification θ ι) : Modification η ι where - app a := Γ.app a ≫ Δ.app a - -end Modification - -/-- Category structure on the oplax natural transformations between OplaxFunctors. -/ -@[simps] -instance category (F G : OplaxFunctor B C) : Category (F ⟶ G) where - Hom := Modification - id := Modification.id - comp := Modification.vcomp - -@[ext] -lemma ext {F G : OplaxFunctor B C} {α β : F ⟶ G} {m n : α ⟶ β} (w : ∀ b, m.app b = n.app b) : - m = n := by - apply Modification.ext - ext - apply w - -@[simp] -lemma Modification.id_app' {X : B} {F G : OplaxFunctor B C} (α : F ⟶ G) : - Modification.app (𝟙 α) X = 𝟙 (α.app X) := rfl - -@[simp] -lemma Modification.comp_app' {X : B} {F G : OplaxFunctor B C} {α β γ : F ⟶ G} - (m : α ⟶ β) (n : β ⟶ γ) : (m ≫ n).app X = m.app X ≫ n.app X := - rfl - -/-- Construct a modification isomorphism between oplax natural transformations -by giving object level isomorphisms, and checking naturality only in the forward direction. --/ -@[simps] -def ModificationIso.ofComponents (app : ∀ a, η.app a ≅ θ.app a) - (naturality : - ∀ {a b} (f : a ⟶ b), - F.map f ◁ (app b).hom ≫ θ.naturality f = η.naturality f ≫ (app a).hom ▷ G.map f) : - η ≅ θ where - hom := { app := fun a => (app a).hom } - inv := - { app := fun a => (app a).inv - naturality := fun {a b} f => by - simpa using congr_arg (fun f => _ ◁ (app b).inv ≫ f ≫ (app a).inv ▷ _) (naturality f).symm } - -end - /-- A structure on an Oplax natural transformation that promotes it to a strong natural transformation. From 5ffbcdc2158ef39e8608723b75b7d2f9e3097a32 Mon Sep 17 00:00:00 2001 From: Wojciech Nawrocki Date: Sat, 26 Oct 2024 23:16:15 +0000 Subject: [PATCH 463/505] doc(CategoryTheory/Reassoc): fix typo (#18207) Co-authored-by: Wojciech Nawrocki --- Mathlib/Tactic/CategoryTheory/Reassoc.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Tactic/CategoryTheory/Reassoc.lean b/Mathlib/Tactic/CategoryTheory/Reassoc.lean index f76eba34bf8a9..7861b38db0151 100644 --- a/Mathlib/Tactic/CategoryTheory/Reassoc.lean +++ b/Mathlib/Tactic/CategoryTheory/Reassoc.lean @@ -63,7 +63,7 @@ that are already right associated. Note that if you want both the lemma and the reassociated lemma to be `simp` lemmas, you should tag the lemma `@[reassoc (attr := simp)]`. The variant `@[simp, reassoc]` on a lemma `F` will tag `F` with `@[simp]`, -but not `F_apply` (this is sometimes useful). +but not `F_assoc` (this is sometimes useful). -/ syntax (name := reassoc) "reassoc" (" (" &"attr" ":=" Parser.Term.attrInstance,* ")")? : attr From 608b89fc0dd5614d6a50ebd8c5a6ff43c05deb87 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sat, 26 Oct 2024 23:35:06 +0000 Subject: [PATCH 464/505] chore: do not depend on CategoryTheory in Module.Injective (#17747) The category-theoretic results can be split between `Mathlib/Algebra/Category/Grp/Injective.lean` and `Mathlib/Algebra/Category/ModuleCat/Injective.lean` instead, with no changes to downstream proofs. Co-authored-by: Junyan Xu --- Mathlib.lean | 1 + .../Category/Grp/EnoughInjectives.lean | 4 +- Mathlib/Algebra/Category/Grp/Injective.lean | 5 +- .../Category/ModuleCat/EnoughInjectives.lean | 34 ++++++++++ .../Algebra/Category/ModuleCat/Injective.lean | 65 +++++++++++-------- Mathlib/Algebra/Module/Injective.lean | 31 +-------- 6 files changed, 80 insertions(+), 60 deletions(-) create mode 100644 Mathlib/Algebra/Category/ModuleCat/EnoughInjectives.lean diff --git a/Mathlib.lean b/Mathlib.lean index 01c450dd99fe5..ed97781e835fb 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -100,6 +100,7 @@ import Mathlib.Algebra.Category.ModuleCat.ChangeOfRings import Mathlib.Algebra.Category.ModuleCat.Colimits import Mathlib.Algebra.Category.ModuleCat.Differentials.Basic import Mathlib.Algebra.Category.ModuleCat.Differentials.Presheaf +import Mathlib.Algebra.Category.ModuleCat.EnoughInjectives import Mathlib.Algebra.Category.ModuleCat.EpiMono import Mathlib.Algebra.Category.ModuleCat.FilteredColimits import Mathlib.Algebra.Category.ModuleCat.Free diff --git a/Mathlib/Algebra/Category/Grp/EnoughInjectives.lean b/Mathlib/Algebra/Category/Grp/EnoughInjectives.lean index a64a093313fe7..b43e1077c56c4 100644 --- a/Mathlib/Algebra/Category/Grp/EnoughInjectives.lean +++ b/Mathlib/Algebra/Category/Grp/EnoughInjectives.lean @@ -7,6 +7,7 @@ Authors: Jujian Zhang, Junyan Xu import Mathlib.Algebra.Module.CharacterModule import Mathlib.Algebra.Category.Grp.EquivalenceGroupAddGroup import Mathlib.Algebra.Category.Grp.EpiMono +import Mathlib.Algebra.Category.Grp.Injective /-! @@ -17,8 +18,7 @@ injective presentation for `A`, hence category of abelian groups has enough inje ## Implementation notes -This file is split from `Mathlib.Algebra.Grp.Injective` is to prevent import loop. -This file's dependency imports `Mathlib.Algebra.Grp.Injective`. +This file is split from `Mathlib.Algebra.Category.Grp.Injective` to prevent import loops. -/ open CategoryTheory diff --git a/Mathlib/Algebra/Category/Grp/Injective.lean b/Mathlib/Algebra/Category/Grp/Injective.lean index 04d0f6d61c4a0..f8ea96a6004ef 100644 --- a/Mathlib/Algebra/Category/Grp/Injective.lean +++ b/Mathlib/Algebra/Category/Grp/Injective.lean @@ -3,9 +3,11 @@ Copyright (c) 2022 Jujian Zhang. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jujian Zhang -/ +import Mathlib.Algebra.Category.ModuleCat.EpiMono import Mathlib.Algebra.Category.Grp.ZModuleEquivalence import Mathlib.Algebra.EuclideanDomain.Int -import Mathlib.Algebra.Module.Injective +import Mathlib.Algebra.Category.ModuleCat.Injective +import Mathlib.CategoryTheory.Preadditive.Injective import Mathlib.RingTheory.PrincipalIdealDomain import Mathlib.Topology.Instances.AddCircle @@ -40,7 +42,6 @@ universe u variable (A : Type u) [AddCommGroup A] - theorem Module.Baer.of_divisible [DivisibleBy A ℤ] : Module.Baer ℤ A := fun I g ↦ by rcases IsPrincipalIdealRing.principal I with ⟨m, rfl⟩ obtain rfl | h0 := eq_or_ne m 0 diff --git a/Mathlib/Algebra/Category/ModuleCat/EnoughInjectives.lean b/Mathlib/Algebra/Category/ModuleCat/EnoughInjectives.lean new file mode 100644 index 0000000000000..d167130dd9438 --- /dev/null +++ b/Mathlib/Algebra/Category/ModuleCat/EnoughInjectives.lean @@ -0,0 +1,34 @@ +/- +Copyright (c) 2023 Jujian Zhang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jujian Zhang +-/ +import Mathlib.Algebra.Category.ModuleCat.ChangeOfRings +import Mathlib.Algebra.Category.ModuleCat.Injective +import Mathlib.Algebra.Category.Grp.EnoughInjectives +import Mathlib.Algebra.Category.Grp.ZModuleEquivalence +import Mathlib.Logic.Equiv.TransferInstance + +/-! +# Category of $R$-modules has enough injectives + +we lift enough injectives of abelian groups to arbitrary $R$-modules by adjoint functors +`restrictScalars ⊣ coextendScalars` +-/ + +open CategoryTheory + +universe v u +variable (R : Type u) [Ring R] + +instance : EnoughInjectives (ModuleCat.{v} ℤ) := + EnoughInjectives.of_equivalence (forget₂ (ModuleCat ℤ) AddCommGrp) + +lemma ModuleCat.enoughInjectives : EnoughInjectives (ModuleCat.{max v u} R) := + EnoughInjectives.of_adjunction (ModuleCat.restrictCoextendScalarsAdj.{max v u} (algebraMap ℤ R)) + +open ModuleCat in +instance [UnivLE.{u,v}] : EnoughInjectives (ModuleCat.{v} R) := + letI := (equivShrink.{v} R).symm.ring + letI := enoughInjectives.{v} (Shrink.{v} R) + EnoughInjectives.of_equivalence (restrictScalars (equivShrink R).symm.ringEquiv.toRingHom) diff --git a/Mathlib/Algebra/Category/ModuleCat/Injective.lean b/Mathlib/Algebra/Category/ModuleCat/Injective.lean index ab64c32a4c6ea..f579799501b6a 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Injective.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Injective.lean @@ -3,36 +3,47 @@ Copyright (c) 2023 Jujian Zhang. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jujian Zhang -/ -import Mathlib.Algebra.Category.ModuleCat.ChangeOfRings -import Mathlib.Algebra.Category.Grp.EnoughInjectives -import Mathlib.Algebra.Category.Grp.ZModuleEquivalence -import Mathlib.Logic.Equiv.TransferInstance +import Mathlib.Algebra.Module.Injective +import Mathlib.CategoryTheory.Preadditive.Injective +import Mathlib.Algebra.Category.ModuleCat.EpiMono /-! -# Category of $R$-modules has enough injectives - -we lift enough injectives of abelian groups to arbitrary $R$-modules by adjoint functors -`restrictScalars ⊣ coextendScalars` - -## Implementation notes -This file is not part of `Algebra/Module/Injective.lean` to prevent import loop: enough-injectives -of abelian groups needs `Algebra/Module/Injective.lean` and this file needs enough-injectives of -abelian groups. +# Injective objects in the category of $R$-modules -/ open CategoryTheory -universe v u -variable (R : Type u) [Ring R] - -instance : EnoughInjectives (ModuleCat.{v} ℤ) := - EnoughInjectives.of_equivalence (forget₂ (ModuleCat ℤ) AddCommGrp) - -lemma ModuleCat.enoughInjectives : EnoughInjectives (ModuleCat.{max v u} R) := - EnoughInjectives.of_adjunction (ModuleCat.restrictCoextendScalarsAdj.{max v u} (algebraMap ℤ R)) - -open ModuleCat in -instance [UnivLE.{u,v}] : EnoughInjectives (ModuleCat.{v} R) := - letI := (equivShrink.{v} R).symm.ring - letI := enoughInjectives.{v} (Shrink.{v} R) - EnoughInjectives.of_equivalence (restrictScalars (equivShrink R).symm.ringEquiv.toRingHom) +universe u v +variable (R : Type u) (M : Type v) [Ring R] [AddCommGroup M] [Module R M] +namespace Module + +theorem injective_object_of_injective_module [inj : Injective R M] : + CategoryTheory.Injective (ModuleCat.of R M) where + factors g f m := + have ⟨l, h⟩ := inj.out f ((ModuleCat.mono_iff_injective f).mp m) g + ⟨l, LinearMap.ext h⟩ + +theorem injective_module_of_injective_object + [inj : CategoryTheory.Injective <| ModuleCat.of R M] : + Module.Injective R M where + out X Y _ _ _ _ f hf g := by + have : CategoryTheory.Mono (ModuleCat.asHom f) := (ModuleCat.mono_iff_injective _).mpr hf + obtain ⟨l, rfl⟩ := inj.factors (ModuleCat.asHom g) (ModuleCat.asHom f) + exact ⟨l, fun _ ↦ rfl⟩ + +theorem injective_iff_injective_object : + Module.Injective R M ↔ + CategoryTheory.Injective (ModuleCat.of R M) := + ⟨fun _ => injective_object_of_injective_module R M, + fun _ => injective_module_of_injective_object R M⟩ + +end Module + + +instance ModuleCat.ulift_injective_of_injective.{v'} + [Small.{v} R] [AddCommGroup M] [Module R M] + [CategoryTheory.Injective <| ModuleCat.of R M] : + CategoryTheory.Injective <| ModuleCat.of R (ULift.{v'} M) := + Module.injective_object_of_injective_module + (inj := Module.ulift_injective_of_injective + (inj := Module.injective_module_of_injective_object _ _)) diff --git a/Mathlib/Algebra/Module/Injective.lean b/Mathlib/Algebra/Module/Injective.lean index 9ed163e6256bf..47a21a20f9c87 100644 --- a/Mathlib/Algebra/Module/Injective.lean +++ b/Mathlib/Algebra/Module/Injective.lean @@ -3,11 +3,10 @@ Copyright (c) 2022 Jujian Zhang. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jujian Zhang -/ -import Mathlib.CategoryTheory.Preadditive.Injective -import Mathlib.Algebra.Category.ModuleCat.EpiMono import Mathlib.RingTheory.Ideal.Basic import Mathlib.LinearAlgebra.LinearPMap import Mathlib.Logic.Equiv.TransferInstance +import Mathlib.Logic.Small.Basic /-! # Injective modules @@ -33,6 +32,7 @@ import Mathlib.Logic.Equiv.TransferInstance -/ +assert_not_exists ModuleCat noncomputable section @@ -57,26 +57,6 @@ map to `Q`, i.e. in the following diagram, if `f` is injective then there is an (f : X →ₗ[R] Y) (_ : Function.Injective f) (g : X →ₗ[R] Q), ∃ h : Y →ₗ[R] Q, ∀ x, h (f x) = g x -theorem Module.injective_object_of_injective_module [inj : Module.Injective R Q] : - CategoryTheory.Injective (ModuleCat.of R Q) where - factors g f m := - have ⟨l, h⟩ := inj.out f ((ModuleCat.mono_iff_injective f).mp m) g - ⟨l, LinearMap.ext h⟩ - -theorem Module.injective_module_of_injective_object - [inj : CategoryTheory.Injective <| ModuleCat.of R Q] : - Module.Injective R Q where - out X Y _ _ _ _ f hf g := by - have : CategoryTheory.Mono (ModuleCat.asHom f) := (ModuleCat.mono_iff_injective _).mpr hf - obtain ⟨l, rfl⟩ := inj.factors (ModuleCat.asHom g) (ModuleCat.asHom f) - exact ⟨l, fun _ ↦ rfl⟩ - -theorem Module.injective_iff_injective_object : - Module.Injective R Q ↔ - CategoryTheory.Injective (ModuleCat.of R Q) := - ⟨fun _ => injective_object_of_injective_module R Q, - fun _ => injective_module_of_injective_object R Q⟩ - /-- An `R`-module `Q` satisfies Baer's criterion if any `R`-linear map from an `Ideal R` extends to an `R`-linear map `R ⟶ Q`-/ def Module.Baer : Prop := @@ -440,13 +420,6 @@ lemma Module.injective_iff_ulift_injective : ⟨Module.ulift_injective_of_injective R, Module.injective_of_ulift_injective R⟩ -instance ModuleCat.ulift_injective_of_injective - [inj : CategoryTheory.Injective <| ModuleCat.of R M] : - CategoryTheory.Injective <| ModuleCat.of R (ULift.{v'} M) := - Module.injective_object_of_injective_module - (inj := Module.ulift_injective_of_injective - (inj := Module.injective_module_of_injective_object (inj := inj))) - end ULift section lifting_property From 824887828166406874a8ef38feefc05597c91a80 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sat, 26 Oct 2024 23:35:08 +0000 Subject: [PATCH 465/505] chore(Submonoid/Inverses): golf (#18259) Golf `leftInvEquiv` by using `IsUnit.unit`. --- Mathlib/GroupTheory/Submonoid/Inverses.lean | 21 ++++++++------------- 1 file changed, 8 insertions(+), 13 deletions(-) diff --git a/Mathlib/GroupTheory/Submonoid/Inverses.lean b/Mathlib/GroupTheory/Submonoid/Inverses.lean index 469f9ece7f56c..a58eb0aad5f62 100644 --- a/Mathlib/GroupTheory/Submonoid/Inverses.lean +++ b/Mathlib/GroupTheory/Submonoid/Inverses.lean @@ -136,20 +136,15 @@ variable (hS : S ≤ IsUnit.submonoid M) `AddEquiv` to `S`."] noncomputable def leftInvEquiv : S.leftInv ≃* S := { S.fromCommLeftInv with - invFun := fun x ↦ by - choose x' hx using hS x.prop - exact ⟨x'.inv, x, hx ▸ x'.inv_val⟩ - left_inv := fun x ↦ - Subtype.eq <| by - dsimp only; generalize_proofs h; rw [← h.choose.mul_left_inj] - conv => rhs; rw [h.choose_spec] - exact h.choose.inv_val.trans (S.mul_fromLeftInv x).symm - right_inv := fun x ↦ by - dsimp only [fromCommLeftInv] + invFun := fun x ↦ ⟨↑(hS x.2).unit⁻¹, x, by simp⟩ + left_inv := by + intro x ext - rw [fromLeftInv_eq_iff] - convert (hS x.prop).choose.inv_val - exact (hS x.prop).choose_spec.symm } + simp [← Units.mul_eq_one_iff_inv_eq] + right_inv := by + rintro ⟨x, hx⟩ + ext + simp [fromLeftInv_eq_iff] } @[to_additive (attr := simp)] theorem fromLeftInv_leftInvEquiv_symm (x : S) : S.fromLeftInv ((S.leftInvEquiv hS).symm x) = x := From e0bd65eedfa4ffe9616fd33085402c213c4cb01f Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Sun, 27 Oct 2024 03:12:09 +0000 Subject: [PATCH 466/505] =?UTF-8?q?feat:=20`=CF=86=20f=20*=20a=20=3D=200`?= =?UTF-8?q?=20if=20`=CF=86=20id=20*=20a=20=3D=200`=20for=20`f=20:=20C(s,?= =?UTF-8?q?=20R)=E2=82=80`=20and=20`=CF=86`=20a=20star=20homomorphism=20(#?= =?UTF-8?q?18128)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- .../ContinuousMap/StoneWeierstrass.lean | 36 +++++++++++++++---- 1 file changed, 30 insertions(+), 6 deletions(-) diff --git a/Mathlib/Topology/ContinuousMap/StoneWeierstrass.lean b/Mathlib/Topology/ContinuousMap/StoneWeierstrass.lean index 54ba6ba6ff1d1..18ef947a9a9db 100644 --- a/Mathlib/Topology/ContinuousMap/StoneWeierstrass.lean +++ b/Mathlib/Topology/ContinuousMap/StoneWeierstrass.lean @@ -491,12 +491,6 @@ open NonUnitalStarAlgebra Submodule namespace ContinuousMap -/- -`set_option maxSynthPendingDepth 2` after https://github.com/leanprover/lean4/pull/4119 -allows use to remove some shortcut instances. --/ -set_option maxSynthPendingDepth 2 - lemma adjoin_id_eq_span_one_union (s : Set 𝕜) : ((StarAlgebra.adjoin 𝕜 {(restrict s (.id 𝕜) : C(s, 𝕜))}) : Set C(s, 𝕜)) = span 𝕜 ({(1 : C(s, 𝕜))} ∪ (adjoin 𝕜 {(restrict s (.id 𝕜) : C(s, 𝕜))})) := by @@ -632,4 +626,34 @@ theorem ContinuousMapZero.induction_on_of_compact {s : Set 𝕜} [Zero s] (h0 : have := (ContinuousMapZero.adjoin_id_dense h0).closure_eq ▸ Set.mem_univ (x := f) exact mem_closure_iff_frequently.mp this |>.mp <| .of_forall h +lemma ContinuousMapZero.nonUnitalStarAlgHom_apply_mul_eq_zero {𝕜 A : Type*} + [RCLike 𝕜] [NonUnitalRing A] [StarRing A] [TopologicalSpace A] [TopologicalSemiring A] + [T2Space A] [Module 𝕜 A] [IsScalarTower 𝕜 A A] {s : Set 𝕜} [Zero s] [CompactSpace s] + (h0 : (0 : s) = (0 : 𝕜)) (φ : C(s, 𝕜)₀ →⋆ₙₐ[𝕜] A) (a : A) (hmul_id : φ (.id h0) * a = 0) + (hmul_star_id : φ (star (.id h0)) * a = 0) (hφ : Continuous φ) (f : C(s, 𝕜)₀) : + φ f * a = 0 := by + induction f using ContinuousMapZero.induction_on_of_compact h0 with + | zero => simp [map_zero] + | id => exact hmul_id + | star_id => exact hmul_star_id + | add _ _ h₁ h₂ => simp only [map_add, add_mul, h₁, h₂, zero_add] + | mul _ _ _ h => simp only [map_mul, mul_assoc, h, mul_zero] + | smul _ _ h => rw [map_smul, smul_mul_assoc, h, smul_zero] + | frequently f h => exact h.mem_of_closed <| isClosed_eq (by fun_prop) continuous_zero + +lemma ContinuousMapZero.mul_nonUnitalStarAlgHom_apply_eq_zero {𝕜 A : Type*} + [RCLike 𝕜] [NonUnitalRing A] [StarRing A] [TopologicalSpace A] [TopologicalSemiring A] + [T2Space A] [Module 𝕜 A] [SMulCommClass 𝕜 A A] {s : Set 𝕜} [Zero s] [CompactSpace s] + (h0 : (0 : s) = (0 : 𝕜)) (φ : C(s, 𝕜)₀ →⋆ₙₐ[𝕜] A) (a : A) (hmul_id : a * φ (.id h0) = 0) + (hmul_star_id : a * φ (star (.id h0)) = 0) (hφ : Continuous φ) (f : C(s, 𝕜)₀) : + a * φ f = 0 := by + induction f using ContinuousMapZero.induction_on_of_compact h0 with + | zero => simp [map_zero] + | id => exact hmul_id + | star_id => exact hmul_star_id + | add _ _ h₁ h₂ => simp only [map_add, mul_add, h₁, h₂, zero_add] + | mul _ _ h _ => simp only [map_mul, ← mul_assoc, h, zero_mul] + | smul _ _ h => rw [map_smul, mul_smul_comm, h, smul_zero] + | frequently f h => exact h.mem_of_closed <| isClosed_eq (by fun_prop) continuous_zero + end ContinuousMapZero From f4ae1ea6c6f6511fec1a5d5f35a433ca2c977592 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Sun, 27 Oct 2024 08:24:19 +0000 Subject: [PATCH 467/505] feat: Int.prime_ofNat_iff (#18240) --- Mathlib/Data/Nat/Prime/Int.lean | 10 ++++++++-- 1 file changed, 8 insertions(+), 2 deletions(-) diff --git a/Mathlib/Data/Nat/Prime/Int.lean b/Mathlib/Data/Nat/Prime/Int.lean index c5f0347d3b6c8..9a6f71bc49518 100644 --- a/Mathlib/Data/Nat/Prime/Int.lean +++ b/Mathlib/Data/Nat/Prime/Int.lean @@ -38,10 +38,16 @@ end Nat namespace Int +-- See note [no_index around OfNat.ofNat] +@[simp] +theorem prime_ofNat_iff {n : ℕ} : + Prime (no_index (OfNat.ofNat n : ℤ)) ↔ Nat.Prime (OfNat.ofNat n) := + Nat.prime_iff_prime_int.symm + theorem prime_two : Prime (2 : ℤ) := - Nat.prime_iff_prime_int.mp Nat.prime_two + prime_ofNat_iff.mpr Nat.prime_two theorem prime_three : Prime (3 : ℤ) := - Nat.prime_iff_prime_int.mp Nat.prime_three + prime_ofNat_iff.mpr Nat.prime_three end Int From 89bb4c76891eeef01b806d7dcfe3f19346ec3c0d Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 27 Oct 2024 08:54:46 +0000 Subject: [PATCH 468/505] refactor: generalise the Ruzsa triangle inequality to non-abelian groups (#18145) The proof works just fine, and the other statements too, up to turning some `*` around. Add supporting lemmas of the form `op '' (s * t) = op '' t * op '' s`. Use the new `#` notation for `Finset.card`. --- .../Algebra/Group/Pointwise/Finset/Basic.lean | 19 ++ .../Algebra/Group/Pointwise/Set/Basic.lean | 5 +- Mathlib/Algebra/Opposites.lean | 6 + Mathlib/Algebra/Pointwise/Stabilizer.lean | 2 +- .../Combinatorics/Additive/DoublingConst.lean | 3 - .../Additive/PluenneckeRuzsa.lean | 191 +++++++++++------- 6 files changed, 147 insertions(+), 79 deletions(-) diff --git a/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean b/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean index b23f7252ab9a3..840cc5164aa40 100644 --- a/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean +++ b/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean @@ -171,6 +171,12 @@ lemma max'_one [LinearOrder α] : (1 : Finset α).max' one_nonempty = 1 := rfl @[to_additive (attr := simp)] lemma min'_one [LinearOrder α] : (1 : Finset α).min' one_nonempty = 1 := rfl +@[to_additive (attr := simp)] +lemma image_op_one [DecidableEq α] : (1 : Finset α).image op = 1 := rfl + +@[to_additive (attr := simp)] +lemma map_op_one : (1 : Finset α).map opEquiv.toEmbedding = 1 := rfl + end One /-! ### Finset negation/inversion -/ @@ -257,6 +263,10 @@ lemma inf'_inv [SemilatticeInf β] {s : Finset α} (hs : s⁻¹.Nonempty) (f : @[to_additive] lemma image_op_inv (s : Finset α) : s⁻¹.image op = (s.image op)⁻¹ := image_comm op_inv +@[to_additive] +lemma map_op_inv (s : Finset α) : s⁻¹.map opEquiv.toEmbedding = (s.map opEquiv.toEmbedding)⁻¹ := by + simp [map_eq_image, image_op_inv] + end Inv open Pointwise @@ -581,6 +591,15 @@ theorem subset_mul {s t : Set α} : theorem image_mul [DecidableEq β] : (s * t).image (f : α → β) = s.image f * t.image f := image_image₂_distrib <| map_mul f +@[to_additive] +lemma image_op_mul (s t : Finset α) : (s * t).image op = t.image op * s.image op := + image_image₂_antidistrib op_mul + +@[to_additive] +lemma map_op_mul (s t : Finset α) : + (s * t).map opEquiv.toEmbedding = t.map opEquiv.toEmbedding * s.map opEquiv.toEmbedding := by + simp [map_eq_image, image_op_mul] + /-- The singleton operation as a `MulHom`. -/ @[to_additive "The singleton operation as an `AddHom`."] def singletonMulHom : α →ₙ* Finset α where diff --git a/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean b/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean index 59cf3b949909d..e62af9d0fc374 100644 --- a/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean +++ b/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean @@ -67,8 +67,7 @@ Because the pointwise action can easily be spelled out in such cases, we give hi nat and int actions. -/ - -open Function +open Function MulOpposite variable {F α β γ : Type*} @@ -132,6 +131,8 @@ noncomputable def singletonOneHom : OneHom α (Set α) where theorem coe_singletonOneHom : (singletonOneHom : α → Set α) = singleton := rfl +@[to_additive] lemma image_op_one : (1 : Set α).image op = 1 := image_singleton + end One /-! ### Set negation/inversion -/ diff --git a/Mathlib/Algebra/Opposites.lean b/Mathlib/Algebra/Opposites.lean index 0fa41a147a2a7..4848307dbf92e 100644 --- a/Mathlib/Algebra/Opposites.lean +++ b/Mathlib/Algebra/Opposites.lean @@ -138,6 +138,12 @@ theorem unop_inj {x y : αᵐᵒᵖ} : unop x = unop y ↔ x = y := attribute [nolint simpComm] AddOpposite.unop_inj +@[to_additive (attr := simp)] lemma «forall» {p : αᵐᵒᵖ → Prop} : (∀ a, p a) ↔ ∀ a, p (op a) := + op_surjective.forall + +@[to_additive (attr := simp)] lemma «exists» {p : αᵐᵒᵖ → Prop} : (∃ a, p a) ↔ ∃ a, p (op a) := + op_surjective.exists + @[to_additive] instance instNontrivial [Nontrivial α] : Nontrivial αᵐᵒᵖ := op_injective.nontrivial @[to_additive] instance instInhabited [Inhabited α] : Inhabited αᵐᵒᵖ := ⟨op default⟩ diff --git a/Mathlib/Algebra/Pointwise/Stabilizer.lean b/Mathlib/Algebra/Pointwise/Stabilizer.lean index de0ffac8e167c..e7a7ecfdf6292 100644 --- a/Mathlib/Algebra/Pointwise/Stabilizer.lean +++ b/Mathlib/Algebra/Pointwise/Stabilizer.lean @@ -74,7 +74,7 @@ lemma stabilizer_subgroup (s : Subgroup G) : stabilizer G (s : Set G) = s := by @[to_additive (attr := simp)] lemma stabilizer_op_subgroup (s : Subgroup G) : stabilizer Gᵐᵒᵖ (s : Set G) = s.op := by simp_rw [SetLike.ext_iff, mem_stabilizer_set] - simp? says simp only [smul_eq_mul_unop, SetLike.mem_coe, Subgroup.mem_op] + simp only [smul_eq_mul_unop, SetLike.mem_coe, Subgroup.mem_op, «forall», unop_op] refine fun a ↦ ⟨fun h ↦ ?_, fun ha b ↦ s.mul_mem_cancel_right ha⟩ simpa only [op_smul_eq_mul, SetLike.mem_coe, one_mul] using (h 1).2 s.one_mem diff --git a/Mathlib/Combinatorics/Additive/DoublingConst.lean b/Mathlib/Combinatorics/Additive/DoublingConst.lean index c4534bdd4b0d4..3212ebdeeb82f 100644 --- a/Mathlib/Combinatorics/Additive/DoublingConst.lean +++ b/Mathlib/Combinatorics/Additive/DoublingConst.lean @@ -5,9 +5,6 @@ Authors: Yaël Dillies -/ import Mathlib.Combinatorics.Additive.PluenneckeRuzsa import Mathlib.Data.Finset.Density -import Mathlib.Tactic.Positivity.Basic -import Mathlib.Tactic.Positivity.Finset -import Mathlib.Tactic.Ring /-! # Doubling and difference constants diff --git a/Mathlib/Combinatorics/Additive/PluenneckeRuzsa.lean b/Mathlib/Combinatorics/Additive/PluenneckeRuzsa.lean index ae04f34d67a52..94e3ef08335bd 100644 --- a/Mathlib/Combinatorics/Additive/PluenneckeRuzsa.lean +++ b/Mathlib/Combinatorics/Additive/PluenneckeRuzsa.lean @@ -8,7 +8,11 @@ import Mathlib.Algebra.Order.Field.Basic import Mathlib.Algebra.Order.Field.Rat import Mathlib.Algebra.Order.Ring.Basic import Mathlib.Combinatorics.Enumerative.DoubleCounting +import Mathlib.Tactic.FieldSimp import Mathlib.Tactic.GCongr +import Mathlib.Tactic.Positivity +import Mathlib.Tactic.Positivity.Finset +import Mathlib.Tactic.Ring /-! # The Plünnecke-Ruzsa inequality @@ -29,52 +33,94 @@ inequality. * [Terrence Tao, Van Vu, *Additive Combinatorics][tao-vu] -/ -open Nat +open MulOpposite Nat open scoped Pointwise namespace Finset +variable {G : Type*} [DecidableEq G] -variable {α : Type*} [CommGroup α] [DecidableEq α] {A B C : Finset α} +section Group +variable [Group G] {A B C : Finset G} + +/-! ### Noncommutative Ruzsa triangle inequality -/ /-- **Ruzsa's triangle inequality**. Division version. -/ @[to_additive "**Ruzsa's triangle inequality**. Subtraction version."] -theorem ruzsa_triangle_inequality_div_div_div (A B C : Finset α) : - (A / C).card * B.card ≤ (A / B).card * (B / C).card := by - rw [← card_product (A / B), ← mul_one ((A / B) ×ˢ (B / C)).card] - refine card_mul_le_card_mul (fun b ac ↦ ac.1 * ac.2 = b) (fun x hx ↦ ?_) +theorem ruzsa_triangle_inequality_div_div_div (A B C : Finset G) : + #(A / C) * #B ≤ #(A / B) * #(C / B) := by + rw [← card_product (A / B), ← mul_one #((A / B) ×ˢ (C / B))] + refine card_mul_le_card_mul (fun b (a, c) ↦ a / c = b) (fun x hx ↦ ?_) fun x _ ↦ card_le_one_iff.2 fun hu hv ↦ ((mem_bipartiteBelow _).1 hu).2.symm.trans ?_ · obtain ⟨a, ha, c, hc, rfl⟩ := mem_div.1 hx - refine card_le_card_of_injOn (fun b ↦ (a / b, b / c)) (fun b hb ↦ ?_) fun b₁ _ b₂ _ h ↦ ?_ + refine card_le_card_of_injOn (fun b ↦ (a / b, c / b)) (fun b hb ↦ ?_) fun b₁ _ b₂ _ h ↦ ?_ · rw [mem_bipartiteAbove] - exact ⟨mk_mem_product (div_mem_div ha hb) (div_mem_div hb hc), div_mul_div_cancel _ _ _⟩ + exact ⟨mk_mem_product (div_mem_div ha hb) (div_mem_div hc hb), div_div_div_cancel_right ..⟩ · exact div_right_injective (Prod.ext_iff.1 h).1 · exact ((mem_bipartiteBelow _).1 hv).2 +/-- **Ruzsa's triangle inequality**. Mulinv-mulinv-mulinv version. -/ +@[to_additive "**Ruzsa's triangle inequality**. Addneg-addneg-addneg version."] +theorem ruzsa_triangle_inequality_mulInv_mulInv_mulInv (A B C : Finset G) : + #(A * C⁻¹) * #B ≤ #(A * B⁻¹) * #(C * B⁻¹) := by + simpa [div_eq_mul_inv] using ruzsa_triangle_inequality_div_div_div A B C + +/-- **Ruzsa's triangle inequality**. Invmul-invmul-invmul version. -/ +@[to_additive "**Ruzsa's triangle inequality**. Negadd-negadd-negadd version."] +theorem ruzsa_triangle_inequality_invMul_invMul_invMul (A B C : Finset G) : + #B * #(A⁻¹ * C) ≤ #(B⁻¹ * A) * #(B⁻¹ * C) := by + simpa [mul_comm, div_eq_mul_inv, ← map_op_mul, ← map_op_inv] using + ruzsa_triangle_inequality_div_div_div (G := Gᵐᵒᵖ) (C.map opEquiv.toEmbedding) + (B.map opEquiv.toEmbedding) (A.map opEquiv.toEmbedding) + + /-- **Ruzsa's triangle inequality**. Div-mul-mul version. -/ @[to_additive "**Ruzsa's triangle inequality**. Sub-add-add version."] -theorem ruzsa_triangle_inequality_div_mul_mul (A B C : Finset α) : - (A / C).card * B.card ≤ (A * B).card * (B * C).card := by - rw [← div_inv_eq_mul, ← card_inv B, ← card_inv (B * C), mul_inv, ← div_eq_mul_inv] - exact ruzsa_triangle_inequality_div_div_div _ _ _ +theorem ruzsa_triangle_inequality_div_mul_mul (A B C : Finset G) : + #(A / C) * #B ≤ #(A * B) * #(C * B) := by + simpa using ruzsa_triangle_inequality_div_div_div A B⁻¹ C + +/-- **Ruzsa's triangle inequality**. Mulinv-mul-mul version. -/ +@[to_additive "**Ruzsa's triangle inequality**. Addneg-add-add version."] +theorem ruzsa_triangle_inequality_mulInv_mul_mul (A B C : Finset G) : + #(A * C⁻¹) * #B ≤ #(A * B) * #(C * B) := by + simpa using ruzsa_triangle_inequality_mulInv_mulInv_mulInv A B⁻¹ C + +/-- **Ruzsa's triangle inequality**. Invmul-mul-mul version. -/ +@[to_additive "**Ruzsa's triangle inequality**. Negadd-add-add version."] +theorem ruzsa_triangle_inequality_invMul_mul_mul (A B C : Finset G) : + #B * #(A⁻¹ * C) ≤ #(B * A) * #(B * C) := by + simpa using ruzsa_triangle_inequality_invMul_invMul_invMul A B⁻¹ C + /-- **Ruzsa's triangle inequality**. Mul-div-mul version. -/ @[to_additive "**Ruzsa's triangle inequality**. Add-sub-add version."] -theorem ruzsa_triangle_inequality_mul_div_mul (A B C : Finset α) : - (A * C).card * B.card ≤ (A / B).card * (B * C).card := by - rw [← div_inv_eq_mul, ← div_inv_eq_mul B] - exact ruzsa_triangle_inequality_div_div_div _ _ _ - -/-- **Ruzsa's triangle inequality**. Mul-mul-div version. -/ -@[to_additive "**Ruzsa's triangle inequality**. Add-add-sub version."] -theorem ruzsa_triangle_inequality_mul_mul_div (A B C : Finset α) : - (A * C).card * B.card ≤ (A * B).card * (B / C).card := by - rw [← div_inv_eq_mul, div_eq_mul_inv B] - exact ruzsa_triangle_inequality_div_mul_mul _ _ _ +theorem ruzsa_triangle_inequality_mul_div_mul (A B C : Finset G) : + #B * #(A * C) ≤ #(B / A) * #(B * C) := by + simpa [div_eq_mul_inv] using ruzsa_triangle_inequality_invMul_mul_mul A⁻¹ B C + +/-- **Ruzsa's triangle inequality**. Mul-mulinv-mul version. -/ +@[to_additive "**Ruzsa's triangle inequality**. Add-addneg-add version."] +theorem ruzsa_triangle_inequality_mul_mulInv_mul (A B C : Finset G) : + #B * #(A * C) ≤ #(B * A⁻¹) * #(B * C) := by + simpa [div_eq_mul_inv] using ruzsa_triangle_inequality_mul_div_mul A B C + +/-- **Ruzsa's triangle inequality**. Mul-mul-invmul version. -/ +@[to_additive "**Ruzsa's triangle inequality**. Add-add-negadd version."] +theorem ruzsa_triangle_inequality_mul_mul_invMul (A B C : Finset G) : + #(A * C) * #B ≤ #(A * B) * #(C⁻¹ * B) := by + simpa using ruzsa_triangle_inequality_mulInv_mul_mul A B C⁻¹ + +end Group + +section CommGroup +variable [CommGroup G] {A B C : Finset G} + +/-! ### Plünnecke-Petridis inequality -/ @[to_additive] -theorem pluennecke_petridis_inequality_mul (C : Finset α) - (hA : ∀ A' ⊆ A, (A * B).card * A'.card ≤ (A' * B).card * A.card) : - (A * B * C).card * A.card ≤ (A * B).card * (A * C).card := by +theorem pluennecke_petridis_inequality_mul (C : Finset G) + (hA : ∀ A' ⊆ A, #(A * B) * #A' ≤ #(A' * B) * #A) : + #(A * B * C) * #A ≤ #(A * B) * #(A * C) := by induction' C using Finset.induction_on with x C _ ih · simp set A' := A ∩ (A * C / {x}) with hA' @@ -88,7 +134,7 @@ theorem pluennecke_petridis_inequality_mul (C : Finset α) exact mul_subset_mul_right inter_subset_right have h₂ : A' * B * {x} ⊆ A * B * {x} := mul_subset_mul_right (mul_subset_mul_right inter_subset_left) - have h₃ : (A * B * C').card ≤ (A * B * C).card + (A * B).card - (A' * B).card := by + have h₃ : #(A * B * C') ≤ #(A * B * C) + #(A * B) - #(A' * B) := by rw [h₁] refine (card_union_le _ _).trans_eq ?_ rw [card_sdiff h₂, ← add_tsub_assoc_of_le (card_le_card h₂), card_mul_singleton, @@ -97,35 +143,34 @@ theorem pluennecke_petridis_inequality_mul (C : Finset α) rw [tsub_mul, add_mul] refine (tsub_le_tsub (add_le_add_right ih _) <| hA _ inter_subset_left).trans_eq ?_ rw [← mul_add, ← mul_tsub, ← hA', hC', insert_eq, mul_union, ← card_mul_singleton A x, ← - card_mul_singleton A' x, add_comm (card _), h₀, + card_mul_singleton A' x, add_comm #_, h₀, eq_tsub_of_add_eq (card_union_add_card_inter _ _)] -/-! ### Sum triangle inequality -/ - +/-! ### Commutative Ruzsa triangle inequality -/ -- Auxiliary lemma for Ruzsa's triangle sum inequality, and the Plünnecke-Ruzsa inequality. @[to_additive] private theorem mul_aux (hA : A.Nonempty) (hAB : A ⊆ B) - (h : ∀ A' ∈ B.powerset.erase ∅, ((A * C).card : ℚ≥0) / ↑A.card ≤ (A' * C).card / ↑A'.card) : - ∀ A' ⊆ A, (A * C).card * A'.card ≤ (A' * C).card * A.card := by + (h : ∀ A' ∈ B.powerset.erase ∅, (#(A * C) : ℚ≥0) / #A ≤ #(A' * C) / #A') : + ∀ A' ⊆ A, #(A * C) * #A' ≤ #(A' * C) * #A := by rintro A' hAA' obtain rfl | hA' := A'.eq_empty_or_nonempty · simp - have hA₀ : (0 : ℚ≥0) < A.card := cast_pos.2 hA.card_pos - have hA₀' : (0 : ℚ≥0) < A'.card := cast_pos.2 hA'.card_pos + have hA₀ : (0 : ℚ≥0) < #A := cast_pos.2 hA.card_pos + have hA₀' : (0 : ℚ≥0) < #A' := cast_pos.2 hA'.card_pos exact mod_cast (div_le_div_iff hA₀ hA₀').1 (h _ <| mem_erase_of_ne_of_mem hA'.ne_empty <| mem_powerset.2 <| hAA'.trans hAB) /-- **Ruzsa's triangle inequality**. Multiplication version. -/ @[to_additive "**Ruzsa's triangle inequality**. Addition version."] -theorem ruzsa_triangle_inequality_mul_mul_mul (A B C : Finset α) : - (A * C).card * B.card ≤ (A * B).card * (B * C).card := by +theorem ruzsa_triangle_inequality_mul_mul_mul (A B C : Finset G) : + #(A * C) * #B ≤ #(A * B) * #(B * C) := by obtain rfl | hB := B.eq_empty_or_nonempty · simp have hB' : B ∈ B.powerset.erase ∅ := mem_erase_of_ne_of_mem hB.ne_empty (mem_powerset_self _) obtain ⟨U, hU, hUA⟩ := - exists_min_image (B.powerset.erase ∅) (fun U ↦ (U * A).card / U.card : _ → ℚ≥0) ⟨B, hB'⟩ + exists_min_image (B.powerset.erase ∅) (fun U ↦ #(U * A) / #U : _ → ℚ≥0) ⟨B, hB'⟩ rw [mem_erase, mem_powerset, ← nonempty_iff_ne_empty] at hU refine cast_le.1 (?_ : (_ : ℚ≥0) ≤ _) push_cast @@ -141,80 +186,80 @@ theorem ruzsa_triangle_inequality_mul_mul_mul (A B C : Finset α) : /-- **Ruzsa's triangle inequality**. Mul-div-div version. -/ @[to_additive "**Ruzsa's triangle inequality**. Add-sub-sub version."] -theorem ruzsa_triangle_inequality_mul_div_div (A B C : Finset α) : - (A * C).card * B.card ≤ (A / B).card * (B / C).card := by +theorem ruzsa_triangle_inequality_mul_div_div (A B C : Finset G) : + #(A * C) * #B ≤ #(A / B) * #(B / C) := by rw [div_eq_mul_inv, ← card_inv B, ← card_inv (B / C), inv_div', div_inv_eq_mul] exact ruzsa_triangle_inequality_mul_mul_mul _ _ _ /-- **Ruzsa's triangle inequality**. Div-mul-div version. -/ @[to_additive "**Ruzsa's triangle inequality**. Sub-add-sub version."] -theorem ruzsa_triangle_inequality_div_mul_div (A B C : Finset α) : - (A / C).card * B.card ≤ (A * B).card * (B / C).card := by +theorem ruzsa_triangle_inequality_div_mul_div (A B C : Finset G) : + #(A / C) * #B ≤ #(A * B) * #(B / C) := by rw [div_eq_mul_inv, div_eq_mul_inv] exact ruzsa_triangle_inequality_mul_mul_mul _ _ _ /-- **Ruzsa's triangle inequality**. Div-div-mul version. -/ @[to_additive "**Ruzsa's triangle inequality**. Sub-sub-add version."] -theorem card_div_mul_le_card_div_mul_card_mul (A B C : Finset α) : - (A / C).card * B.card ≤ (A / B).card * (B * C).card := by +theorem card_div_mul_le_card_div_mul_card_mul (A B C : Finset G) : + #(A / C) * #B ≤ #(A / B) * #(B * C) := by rw [← div_inv_eq_mul, div_eq_mul_inv] exact ruzsa_triangle_inequality_mul_div_div _ _ _ -- Auxiliary lemma towards the Plünnecke-Ruzsa inequality @[to_additive] -private lemma card_mul_pow_le (hAB : ∀ A' ⊆ A, (A * B).card * A'.card ≤ (A' * B).card * A.card) - (n : ℕ) : (A * B ^ n).card ≤ ((A * B).card / A.card : ℚ≥0) ^ n * A.card := by +private lemma card_mul_pow_le (hAB : ∀ A' ⊆ A, #(A * B) * #A' ≤ #(A' * B) * #A) (n : ℕ) : + #(A * B ^ n) ≤ (#(A * B) / #A : ℚ≥0) ^ n * #A := by obtain rfl | hA := A.eq_empty_or_nonempty · simp induction' n with n ih · simp - rw [_root_.pow_succ', ← mul_assoc, _root_.pow_succ', @mul_assoc ℚ≥0, ← mul_div_right_comm, - le_div_iff₀, ← cast_mul] - swap - · exact cast_pos.2 hA.card_pos - refine (Nat.cast_le.2 <| pluennecke_petridis_inequality_mul _ hAB).trans ?_ - rw [cast_mul] - gcongr + refine le_of_mul_le_mul_right ?_ (by positivity : (0 : ℚ≥0) < #A) + calc + ((#(A * B ^ (n + 1))) * #A : ℚ≥0) + = #(A * B * B ^ n) * #A := by rw [_root_.pow_succ', ← mul_assoc] + _ ≤ #(A * B) * #(A * B ^ n) := mod_cast pluennecke_petridis_inequality_mul _ hAB + _ ≤ #(A * B) * ((#(A * B) / #A) ^ n * #A) := by gcongr + _ = (#(A * B) / #A) ^ (n + 1) * #A * #A := by field_simp; ring /-- The **Plünnecke-Ruzsa inequality**. Multiplication version. Note that this is genuinely harder than the division version because we cannot use a double counting argument. -/ @[to_additive "The **Plünnecke-Ruzsa inequality**. Addition version. Note that this is genuinely harder than the subtraction version because we cannot use a double counting argument."] -theorem pluennecke_ruzsa_inequality_pow_div_pow_mul (hA : A.Nonempty) (B : Finset α) (m n : ℕ) : - ((B ^ m / B ^ n).card) ≤ ((A * B).card / A.card : ℚ≥0) ^ (m + n) * A.card := by +theorem pluennecke_ruzsa_inequality_pow_div_pow_mul (hA : A.Nonempty) (B : Finset G) (m n : ℕ) : + #(B ^ m / B ^ n) ≤ (#(A * B) / #A : ℚ≥0) ^ (m + n) * #A := by have hA' : A ∈ A.powerset.erase ∅ := mem_erase_of_ne_of_mem hA.ne_empty (mem_powerset_self _) - obtain ⟨C, hC, hCA⟩ := - exists_min_image (A.powerset.erase ∅) (fun C ↦ (C * B).card / C.card : _ → ℚ≥0) ⟨A, hA'⟩ + obtain ⟨C, hC, hCmin⟩ := + exists_min_image (A.powerset.erase ∅) (fun C ↦ #(C * B) / #C : _ → ℚ≥0) ⟨A, hA'⟩ rw [mem_erase, mem_powerset, ← nonempty_iff_ne_empty] at hC - refine (_root_.mul_le_mul_right <| cast_pos.2 hC.1.card_pos).1 ?_ - norm_cast - refine (Nat.cast_le.2 <| ruzsa_triangle_inequality_div_mul_mul _ _ _).trans ?_ - push_cast - rw [mul_comm _ C] - refine (mul_le_mul (card_mul_pow_le (mul_aux hC.1 hC.2 hCA) _) - (card_mul_pow_le (mul_aux hC.1 hC.2 hCA) _) (zero_le _) (zero_le _)).trans ?_ - rw [mul_mul_mul_comm, ← pow_add, ← mul_assoc] - gcongr ((?_ ^ _) * Nat.cast ?_) * _ - · exact hCA _ hA' - · exact card_le_card hC.2 + obtain ⟨hC, hCA⟩ := hC + refine le_of_mul_le_mul_right ?_ (by positivity : (0 : ℚ≥0) < #C) + calc + (#(B ^ m / B ^ n) * #C : ℚ≥0) + ≤ #(B ^ m * C) * #(B ^ n * C) := mod_cast ruzsa_triangle_inequality_div_mul_mul .. + _ = #(C * B ^ m) * #(C * B ^ n) := by simp_rw [mul_comm] + _ ≤ ((#(C * B) / #C) ^ m * #C) * ((#(C * B) / #C : ℚ≥0) ^ n * #C) := by + gcongr <;> exact card_mul_pow_le (mul_aux hC hCA hCmin) _ + _ = (#(C * B) / #C) ^ (m + n) * #C * #C := by ring + _ ≤ (#(A * B) / #A) ^ (m + n) * #A * #C := by gcongr (?_ ^ _) * #?_ * _; exact hCmin _ hA' /-- The **Plünnecke-Ruzsa inequality**. Division version. -/ @[to_additive "The **Plünnecke-Ruzsa inequality**. Subtraction version."] -theorem pluennecke_ruzsa_inequality_pow_div_pow_div (hA : A.Nonempty) (B : Finset α) (m n : ℕ) : - (B ^ m / B ^ n).card ≤ ((A / B).card / A.card : ℚ≥0) ^ (m + n) * A.card := by +theorem pluennecke_ruzsa_inequality_pow_div_pow_div (hA : A.Nonempty) (B : Finset G) (m n : ℕ) : + #(B ^ m / B ^ n) ≤ (#(A / B) / #A : ℚ≥0) ^ (m + n) * #A := by rw [← card_inv, inv_div', ← inv_pow, ← inv_pow, div_eq_mul_inv A] exact pluennecke_ruzsa_inequality_pow_div_pow_mul hA _ _ _ /-- Special case of the **Plünnecke-Ruzsa inequality**. Multiplication version. -/ @[to_additive "Special case of the **Plünnecke-Ruzsa inequality**. Addition version."] -theorem pluennecke_ruzsa_inequality_pow_mul (hA : A.Nonempty) (B : Finset α) (n : ℕ) : - (B ^ n).card ≤ ((A * B).card / A.card : ℚ≥0) ^ n * A.card := by +theorem pluennecke_ruzsa_inequality_pow_mul (hA : A.Nonempty) (B : Finset G) (n : ℕ) : + #(B ^ n) ≤ (#(A * B) / #A : ℚ≥0) ^ n * #A := by simpa only [_root_.pow_zero, div_one] using pluennecke_ruzsa_inequality_pow_div_pow_mul hA _ _ 0 /-- Special case of the **Plünnecke-Ruzsa inequality**. Division version. -/ @[to_additive "Special case of the **Plünnecke-Ruzsa inequality**. Subtraction version."] -theorem pluennecke_ruzsa_inequality_pow_div (hA : A.Nonempty) (B : Finset α) (n : ℕ) : - (B ^ n).card ≤ ((A / B).card / A.card : ℚ≥0) ^ n * A.card := by +theorem pluennecke_ruzsa_inequality_pow_div (hA : A.Nonempty) (B : Finset G) (n : ℕ) : + #(B ^ n) ≤ (#(A / B) / #A : ℚ≥0) ^ n * #A := by simpa only [_root_.pow_zero, div_one] using pluennecke_ruzsa_inequality_pow_div_pow_div hA _ _ 0 +end CommGroup end Finset From 5f6de6dd631a9c004553e88b069c9c3908af1299 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Sun, 27 Oct 2024 09:50:50 +0000 Subject: [PATCH 469/505] chore: deprecate Cardinal.mk'_def (#18205) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Violeta Hernández --- Mathlib/SetTheory/Cardinal/Arithmetic.lean | 6 +++--- Mathlib/SetTheory/Cardinal/Basic.lean | 7 +++---- Mathlib/SetTheory/Cardinal/Cofinality.lean | 7 ++++--- Mathlib/SetTheory/Ordinal/Basic.lean | 5 +---- 4 files changed, 11 insertions(+), 14 deletions(-) diff --git a/Mathlib/SetTheory/Cardinal/Arithmetic.lean b/Mathlib/SetTheory/Cardinal/Arithmetic.lean index a002340504602..0bdbc4b83972d 100644 --- a/Mathlib/SetTheory/Cardinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Cardinal/Arithmetic.lean @@ -43,7 +43,7 @@ section mul theorem mul_eq_self {c : Cardinal} (h : ℵ₀ ≤ c) : c * c = c := by refine le_antisymm ?_ (by simpa only [mul_one] using mul_le_mul_left' (one_le_aleph0.trans h) c) -- the only nontrivial part is `c * c ≤ c`. We prove it inductively. - refine Acc.recOn (Cardinal.lt_wf.apply c) (fun c _ => Quotient.inductionOn c fun α IH ol => ?_) h + refine Acc.recOn (Cardinal.lt_wf.apply c) (fun c _ => Cardinal.inductionOn c fun α IH ol => ?_) h -- consider the minimal well-order `r` on `α` (a type with cardinality `c`). rcases ord_eq α with ⟨r, wo, e⟩ classical @@ -80,10 +80,10 @@ theorem mul_eq_self {c : Cardinal} (h : ℵ₀ ≤ c) : c * c = c := by apply @irrefl _ r cases' lt_or_le (card (succ (typein (· < ·) (g p)))) ℵ₀ with qo qo · exact (mul_lt_aleph0 qo qo).trans_le ol - · suffices (succ (typein LT.lt (g p))).card < ⟦α⟧ from (IH _ this qo).trans_lt this + · suffices (succ (typein LT.lt (g p))).card < #α from (IH _ this qo).trans_lt this rw [← lt_ord] apply (isLimit_ord ol).2 - rw [mk'_def, e] + rw [e] apply typein_lt_type /-- If `α` and `β` are infinite types, then the cardinality of `α × β` is the maximum diff --git a/Mathlib/SetTheory/Cardinal/Basic.lean b/Mathlib/SetTheory/Cardinal/Basic.lean index d330b1ffc5407..7443925cdf82b 100644 --- a/Mathlib/SetTheory/Cardinal/Basic.lean +++ b/Mathlib/SetTheory/Cardinal/Basic.lean @@ -141,7 +141,8 @@ theorem induction_on_pi {ι : Type u} {p : (ι → Cardinal.{v}) → Prop} protected theorem eq : #α = #β ↔ Nonempty (α ≃ β) := Quotient.eq' -@[simp] +/-- Avoid using `Quotient.mk` to construct a `Cardinal` directly -/ +@[deprecated (since := "2024-10-24")] theorem mk'_def (α : Type u) : @Eq Cardinal ⟦α⟧ #α := rfl @@ -280,9 +281,7 @@ lemma mk_preimage_down {s : Set α} : #(ULift.down.{v} ⁻¹' s) = lift.{v} (#s) exact Equiv.ofBijective f this theorem out_embedding {c c' : Cardinal} : c ≤ c' ↔ Nonempty (c.out ↪ c'.out) := by - trans - · rw [← Quotient.out_eq c, ← Quotient.out_eq c'] - · rw [mk'_def, mk'_def, le_def] + conv_lhs => rw [← Cardinal.mk_out c, ← Cardinal.mk_out c', le_def] theorem lift_mk_le {α : Type v} {β : Type w} : lift.{max u w} #α ≤ lift.{max u v} #β ↔ Nonempty (α ↪ β) := diff --git a/Mathlib/SetTheory/Cardinal/Cofinality.lean b/Mathlib/SetTheory/Cardinal/Cofinality.lean index 995e87d30575d..3e7d59f172e0d 100644 --- a/Mathlib/SetTheory/Cardinal/Cofinality.lean +++ b/Mathlib/SetTheory/Cardinal/Cofinality.lean @@ -908,7 +908,8 @@ theorem isRegular_succ {c : Cardinal.{u}} (h : ℵ₀ ≤ c) : IsRegular (succ c ⟨h.trans (le_succ c), succ_le_of_lt (by - cases' Quotient.exists_rep (@succ Cardinal _ _ c) with α αe; simp only [mk'_def] at αe + have αe := Cardinal.mk_out (succ c) + set α := (succ c).out rcases ord_eq α with ⟨r, wo, re⟩ have := isLimit_ord (h.trans (le_succ _)) rw [← αe, re] at this ⊢ @@ -1167,10 +1168,10 @@ theorem univ_inaccessible : IsInaccessible univ.{u, v} := apply lift_lt_univ' theorem lt_power_cof {c : Cardinal.{u}} : ℵ₀ ≤ c → c < (c^cof c.ord) := - Quotient.inductionOn c fun α h => by + Cardinal.inductionOn c fun α h => by rcases ord_eq α with ⟨r, wo, re⟩ have := isLimit_ord h - rw [mk'_def, re] at this ⊢ + rw [re] at this ⊢ rcases cof_eq' r this with ⟨S, H, Se⟩ have := sum_lt_prod (fun a : S => #{ x // r x a }) (fun _ => #α) fun i => ?_ · simp only [Cardinal.prod_const, Cardinal.lift_id, ← Se, ← mk_sigma, power_def] at this ⊢ diff --git a/Mathlib/SetTheory/Ordinal/Basic.lean b/Mathlib/SetTheory/Ordinal/Basic.lean index abdf8207a1c0c..c4925eb265a0c 100644 --- a/Mathlib/SetTheory/Ordinal/Basic.lean +++ b/Mathlib/SetTheory/Ordinal/Basic.lean @@ -1135,10 +1135,7 @@ theorem lt_ord {c o} : o < ord c ↔ o.card < c := @[simp] theorem card_ord (c) : (ord c).card = c := - Quotient.inductionOn c fun α => by - let ⟨r, _, e⟩ := ord_eq α - -- Porting note: cardinal.mk_def is now Cardinal.mk'_def, not sure why - simp only [mk'_def, e, card_type] + c.inductionOn fun α ↦ let ⟨r, _, e⟩ := ord_eq α; e ▸ card_type r theorem card_surjective : Function.Surjective card := fun c ↦ ⟨_, card_ord c⟩ From 7e4aaba2363bbceb287b74f286f51a50e868f818 Mon Sep 17 00:00:00 2001 From: Dagur Asgeirsson Date: Sun, 27 Oct 2024 10:17:43 +0000 Subject: [PATCH 470/505] feat(CategoryTheory/Monoidal): Day's reflection theorem (#15894) --- Mathlib.lean | 1 + Mathlib/CategoryTheory/Monad/Basic.lean | 28 +++ .../Monoidal/Braided/Reflection.lean | 233 ++++++++++++++++++ Mathlib/CategoryTheory/Yoneda.lean | 26 ++ docs/references.bib | 14 ++ 5 files changed, 302 insertions(+) create mode 100644 Mathlib/CategoryTheory/Monoidal/Braided/Reflection.lean diff --git a/Mathlib.lean b/Mathlib.lean index ed97781e835fb..eb696d6f93a41 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1819,6 +1819,7 @@ import Mathlib.CategoryTheory.Monoidal.Bimod import Mathlib.CategoryTheory.Monoidal.Bimon_ import Mathlib.CategoryTheory.Monoidal.Braided.Basic import Mathlib.CategoryTheory.Monoidal.Braided.Opposite +import Mathlib.CategoryTheory.Monoidal.Braided.Reflection import Mathlib.CategoryTheory.Monoidal.Cartesian.Comon_ import Mathlib.CategoryTheory.Monoidal.Category import Mathlib.CategoryTheory.Monoidal.Center diff --git a/Mathlib/CategoryTheory/Monad/Basic.lean b/Mathlib/CategoryTheory/Monad/Basic.lean index 6a82f3d208edd..9fe82e5aa7b80 100644 --- a/Mathlib/CategoryTheory/Monad/Basic.lean +++ b/Mathlib/CategoryTheory/Monad/Basic.lean @@ -353,4 +353,32 @@ def transport {F : C ⥤ C} (T : Comonad C) (i : (T : C ⥤ C) ≅ F) : Comonad end Comonad +namespace Monad + +lemma map_unit_app (T : Monad C) (X : C) [IsIso T.μ] : + T.map (T.η.app X) = T.η.app (T.obj X) := by + simp [← cancel_mono (T.μ.app _)] + +lemma isSplitMono_iff_isIso_unit (T : Monad C) (X : C) [IsIso T.μ] : + IsSplitMono (T.η.app X) ↔ IsIso (T.η.app X) := by + refine ⟨fun _ ↦ ⟨retraction (T.η.app X), by simp, ?_⟩, fun _ ↦ inferInstance⟩ + erw [← map_id, ← IsSplitMono.id (T.η.app X), map_comp, T.map_unit_app X, T.η.naturality] + rfl + +end Monad + +namespace Comonad + +lemma map_counit_app (T : Comonad C) (X : C) [IsIso T.δ] : + T.map (T.ε.app X) = T.ε.app (T.obj X) := by + simp [← cancel_epi (T.δ.app _)] + +lemma isSplitEpi_iff_isIso_counit (T : Comonad C) (X : C) [IsIso T.δ] : + IsSplitEpi (T.ε.app X) ↔ IsIso (T.ε.app X) := by + refine ⟨fun _ ↦ ⟨section_ (T.ε.app X), ?_, by simp⟩, fun _ ↦ inferInstance⟩ + erw [← map_id, ← IsSplitEpi.id (T.ε.app X), map_comp, T.map_counit_app X, T.ε.naturality] + rfl + +end Comonad + end CategoryTheory diff --git a/Mathlib/CategoryTheory/Monoidal/Braided/Reflection.lean b/Mathlib/CategoryTheory/Monoidal/Braided/Reflection.lean new file mode 100644 index 0000000000000..665358fbb9504 --- /dev/null +++ b/Mathlib/CategoryTheory/Monoidal/Braided/Reflection.lean @@ -0,0 +1,233 @@ +/- +Copyright (c) 2024 Dagur Asgeirsson. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Dagur Asgeirsson +-/ +import Mathlib.CategoryTheory.Adjunction.Restrict +import Mathlib.CategoryTheory.Closed.Monoidal +import Mathlib.CategoryTheory.Monad.Adjunction +import Mathlib.CategoryTheory.Monoidal.Braided.Basic +import Mathlib.Tactic.TFAE +/-! + +# Day's reflection theorem + +Let `D` be a symmetric monoidal closed category and let `C` be a reflective subcategory. Day's +reflection theorem proves the equivalence of four conditions, which are all of the form that a +map obtained by acting on the unit of the reflective adjunction, with the internal hom and +tensor functors, is an isomorphism. + +Suppose that `C` is itself monoidal and that the reflector is a monoidal functor. Then we can +apply Day's reflection theorem to prove that `C` is also closed monoidal. + +## References + +- We follow the proof on nLab, see https://ncatlab.org/nlab/show/Day%27s+reflection+theorem. +- The original paper is [day1972] *A reflection theorem for closed categories*, by Day, 1972. +-/ + +open CategoryTheory Category MonoidalCategory MonoidalClosed BraidedCategory Functor + +namespace CategoryTheory.Monoidal.Reflective + +variable {C D : Type*} [Category C] [Category D] + +variable [MonoidalCategory D] [SymmetricCategory D] [MonoidalClosed D] + +section +variable {R : C ⥤ D} [R.Faithful] [R.Full] {L : D ⥤ C} (adj : L ⊣ R) + +/-- The uncurried retraction of the unit in the proof of `4 → 1` in `day_reflection` below. -/ +private noncomputable def adjRetractionAux + (c : C) (d : D) [IsIso (L.map (adj.unit.app ((ihom d).obj (R.obj c)) ⊗ adj.unit.app d))] : + d ⊗ ((L ⋙ R).obj ((ihom d).obj (R.obj c))) ⟶ (R.obj c) := + (β_ _ _).hom ≫ (_ ◁ adj.unit.app _) ≫ adj.unit.app _ ≫ + R.map (inv (L.map (adj.unit.app _ ⊗ adj.unit.app _))) ≫ (L ⋙ R).map (β_ _ _).hom ≫ + (L ⋙ R).map ((ihom.ev _).app _) ≫ inv (adj.unit.app _) + +/-- The retraction of the unit in the proof of `4 → 1` in `day_reflection` below. -/ +private noncomputable def adjRetraction (c : C) (d : D) + [IsIso (L.map (adj.unit.app ((ihom d).obj (R.obj c)) ⊗ adj.unit.app d))] : + (L ⋙ R).obj ((ihom d).obj (R.obj c)) ⟶ ((ihom d).obj (R.obj c)) := + curry <| adjRetractionAux adj c d + +private lemma adjRetraction_is_retraction (c : C) (d : D) + [IsIso (L.map (adj.unit.app ((ihom d).obj (R.obj c)) ⊗ adj.unit.app d))] : + adj.unit.app ((ihom d).obj (R.obj c)) ≫ adjRetraction adj c d = 𝟙 _ := by + suffices (_ ◁ adj.unit.app _) ≫ adjRetractionAux adj c d = (ihom.ev _).app _ by + simp only [id_obj, comp_obj, adjRetraction, ← curry_natural_left, this] + simp [curry_eq] + simp only [id_obj, comp_obj, adjRetractionAux, Functor.map_inv, Functor.comp_map, + braiding_naturality_right_assoc] + slice_lhs 2 3 => + simp only [← id_tensorHom, ← tensorHom_id, ← tensor_comp, Category.id_comp, Category.comp_id] + slice_lhs 2 4 => + rw [← adj.unit_naturality_assoc] + simp + +attribute [local simp] Adjunction.homEquiv_unit Adjunction.homEquiv_counit + +/-- +Day's reflection theorem. + +Let `D` be a symmetric monoidal closed category and let `C` be a reflective subcategory. Denote by +`R : C ⥤ D` the inclusion functor and by `L : D ⥤ C` the reflector. Let `u` denote the unit of the +adjunction `L ⊣ R`. Denote the internal hom by `[-,-]`. The following are equivalent: + +1. `u : [d, Rc] ⟶ RL[d, Rc]` is an isomorphism, +2. `[u, 𝟙] : [RLd, Rc] ⟶ [d, Rc]` is an isomorphism, +3. `L(u ▷ d') : L(d ⊗ d') ⟶ L(RLd ⊗ d')` is an isomorphism, +4. `L(u ⊗ u) : L(d ⊗ d') ⟶ L(RLd ⊗ RLd')` is an isomorphism, + +where `c, d, d'` are arbitrary objects of `C`/`D`, quantified over separately in each condition. +-/ +theorem isIso_tfae : List.TFAE + [ ∀ (c : C) (d : D), IsIso (adj.unit.app ((ihom d).obj (R.obj c))) + , ∀ (c : C) (d : D), IsIso ((pre (adj.unit.app d)).app (R.obj c)) + , ∀ (d d' : D), IsIso (L.map ((adj.unit.app d) ▷ d')) + , ∀ (d d' : D), IsIso (L.map ((adj.unit.app d) ⊗ (adj.unit.app d')))] := by + tfae_have 3 → 4 + · intro h + -- We can commute the tensor product in the condition that `L.map ((adj.unit.app d) ▷ d')` is + -- an isomorphism: + have h' : ∀ d d', IsIso (L.map (d ◁ (adj.unit.app d'))) := by + intro d d' + have := braiding_naturality (𝟙 d) (adj.unit.app d') + rw [← Iso.eq_comp_inv, id_tensorHom] at this + rw [this] + simp only [map_comp, id_obj, comp_obj, tensorHom_id, Category.assoc] + infer_instance + intro d d' + -- We then write the tensor product of the two units as the composition of the whiskered units, + -- and conclude. + have : (adj.unit.app d) ⊗ (adj.unit.app d') = + (adj.unit.app d ▷ d') ≫ (((L ⋙ R).obj _) ◁ adj.unit.app d') := by + simp [← tensorHom_id, ← id_tensorHom, ← tensor_comp] + rw [this, map_comp] + infer_instance + tfae_have 4 → 1 + · intros + -- It is enough to show that the unit is a split monomorphism, and the retraction is given + -- by `adjRetraction` above. + let _ : Reflective R := { L := L, adj := adj } + have : IsIso adj.toMonad.μ := μ_iso_of_reflective (R := R) + erw [← adj.toMonad.isSplitMono_iff_isIso_unit] + exact ⟨⟨adjRetraction adj _ _, adjRetraction_is_retraction adj _ _⟩⟩ + tfae_have 1 → 3 + · intro h d d' + rw [isIso_iff_isIso_coyoneda_map] + intro c + -- `w₁, w₃, w₄` are the three stacked commutative squares in the proof on nLab: + have w₁ : (coyoneda.map (L.map (adj.unit.app d ▷ d')).op).app c = (adj.homEquiv _ _).symm ∘ + (coyoneda.map (adj.unit.app d ▷ d').op).app (R.obj c) ∘ adj.homEquiv _ _ := by ext; simp + rw [w₁, isIso_iff_bijective] + simp only [comp_obj, coyoneda_obj_obj, id_obj, EquivLike.comp_bijective, + EquivLike.bijective_comp] + -- We commute the tensor product using the auxiliary commutative square `w₂`. + have w₂ : ((coyoneda.map (adj.unit.app d ▷ d').op).app (R.obj c)) = + ((yoneda.obj (R.obj c)).mapIso (β_ _ _)).hom ∘ + ((coyoneda.map (d' ◁ adj.unit.app d).op).app (R.obj c)) ∘ + ((yoneda.obj (R.obj c)).mapIso (β_ _ _)).hom := by ext; simp + rw [w₂, ← types_comp, ← types_comp, ← isIso_iff_bijective] + refine IsIso.comp_isIso' (IsIso.comp_isIso' inferInstance ?_) inferInstance + have w₃ : ((coyoneda.map (d' ◁ adj.unit.app d).op).app (R.obj c)) = + ((ihom.adjunction d').homEquiv _ _).symm ∘ + ((coyoneda.map (adj.unit.app _).op).app _) ∘ (ihom.adjunction d').homEquiv _ _ := by + ext + simp only [id_obj, op_tensorObj, coyoneda_obj_obj, unop_tensorObj, comp_obj, + coyoneda_map_app, Quiver.Hom.unop_op, Function.comp_apply, + Adjunction.homEquiv_unit, Adjunction.homEquiv_counit] + simp + rw [w₃, isIso_iff_bijective] + simp only [comp_obj, op_tensorObj, coyoneda_obj_obj, unop_tensorObj, id_obj, + yoneda_obj_obj, tensorLeft_obj, EquivLike.comp_bijective, EquivLike.bijective_comp] + have w₄ : (coyoneda.map (adj.unit.app d).op).app ((ihom d').obj (R.obj c)) ≫ + (coyoneda.obj ⟨d⟩).map (adj.unit.app ((ihom d').obj (R.obj c))) = + (coyoneda.obj ⟨(L ⋙ R).obj d⟩).map (adj.unit.app ((ihom d').obj (R.obj c))) ≫ + (coyoneda.map (adj.unit.app d).op).app _ := by simp + rw [← isIso_iff_bijective] + suffices IsIso ((coyoneda.map (adj.unit.app d).op).app ((ihom d').obj (R.obj c)) ≫ + (coyoneda.obj ⟨d⟩).map (adj.unit.app ((ihom d').obj (R.obj c)))) from + IsIso.of_isIso_comp_right _ ((coyoneda.obj ⟨d⟩).map (adj.unit.app ((ihom d').obj (R.obj c)))) + rw [w₄] + refine IsIso.comp_isIso' inferInstance ?_ + constructor + -- We give the inverse of the bottom map in the stack of commutative squares: + refine ⟨fun f ↦ R.map ((adj.homEquiv _ _).symm f), ?_, by ext; simp⟩ + ext f + simp only [comp_obj, coyoneda_obj_obj, id_obj, Adjunction.homEquiv_counit, + map_comp, types_comp_apply, coyoneda_map_app, Quiver.Hom.unop_op, Category.assoc, + types_id_apply] + have : f = R.map (R.preimage f) := by simp + rw [this] + simp [← map_comp, ← map_comp_assoc, -map_preimage] + tfae_have 2 ↔ 3 + · conv => lhs; intro c d; rw [isIso_iff_isIso_yoneda_map] + conv => rhs; intro d d'; rw [isIso_iff_isIso_coyoneda_map] + -- bring the quantifiers out of the `↔`: + rw [forall_swap]; apply forall_congr'; intro d + rw [forall_swap]; apply forall₂_congr; intro d' c + -- `w₁, w₂,` are the two stacked commutative squares in the proof on nLab: + have w₁ : ((coyoneda.map (L.map (adj.unit.app d ▷ d')).op).app c) = + (adj.homEquiv _ _).symm ∘ + (coyoneda.map (adj.unit.app d ▷ d').op).app (R.obj c) ∘ + (adj.homEquiv _ _) := by ext; simp + have w₂ : ((yoneda.map ((pre (adj.unit.app d)).app (R.obj c))).app ⟨d'⟩) = + ((ihom.adjunction d).homEquiv _ _) ∘ + ((coyoneda.map (adj.unit.app d ▷ d').op).app (R.obj c)) ∘ + ((ihom.adjunction ((L ⋙ R).obj d)).homEquiv _ _).symm := by + rw [← Function.comp_assoc, ((ihom.adjunction ((L ⋙ R).obj d)).homEquiv _ _).eq_comp_symm] + ext + simp only [id_obj, yoneda_obj_obj, comp_obj, Function.comp_apply, + yoneda_map_app, op_tensorObj, coyoneda_obj_obj, unop_tensorObj, op_whiskerRight, + coyoneda_map_app, unop_whiskerRight, Quiver.Hom.unop_op] + rw [Adjunction.homEquiv_unit, Adjunction.homEquiv_unit] + simp + rw [w₂, w₁, isIso_iff_bijective, isIso_iff_bijective] + simp + tfae_finish + +end + +section +variable [MonoidalCategory C] +variable {L : MonoidalFunctor D C} {R : C ⥤ D} [R.Faithful] [R.Full] (adj : L.toFunctor ⊣ R) + +instance (d d' : D) : IsIso (L.map ((adj.unit.app d) ⊗ (adj.unit.app d'))) := by + have := L.μ_natural (adj.unit.app d) (adj.unit.app d') + change _ = (asIso _).hom ≫ _ at this + rw [← Iso.inv_comp_eq] at this + rw [← this] + infer_instance + +instance (c : C) (d : D) : IsIso (adj.unit.app ((ihom d).obj (R.obj c))) := by + revert c d + rw [((isIso_tfae adj).out 0 3:)] + intro d d' + infer_instance + +/-- Auxiliary definition for `monoidalClosed`. -/ +noncomputable def closed (c : C) : Closed c where + rightAdj := R ⋙ (ihom (R.obj c)) ⋙ L.toFunctor + adj := by + refine ((ihom.adjunction (R.obj c)).comp adj).restrictFullyFaithful + (FullyFaithful.ofFullyFaithful R) + (FullyFaithful.id _) ?_ ?_ + · refine NatIso.ofComponents (fun _ ↦ ?_) (fun _ ↦ ?_) + · exact (asIso (L.μ _ _)).symm ≪≫ asIso ((adj.counit.app _) ⊗ (adj.counit.app _)) + · simp only [comp_obj, id_obj, Functor.comp_map, tensorLeft_map, Iso.trans_hom, Iso.symm_hom, + asIso_inv, asIso_hom, Functor.id_map, Category.assoc, IsIso.eq_inv_comp] + rw [← L.μ_natural_right_assoc] + simp [← id_tensorHom, ← tensor_comp] + · exact NatIso.ofComponents (fun _ ↦ asIso (adj.unit.app ((ihom _).obj _))) + +/-- +Given a reflective functor `R : C ⥤ D` with a monoidal left adjoint, such that `D` is symmetric +monoidal closed, then `C` is monoidal closed. +-/ +noncomputable def monoidalClosed : MonoidalClosed C where + closed c := closed adj c + +end + +end CategoryTheory.Monoidal.Reflective diff --git a/Mathlib/CategoryTheory/Yoneda.lean b/Mathlib/CategoryTheory/Yoneda.lean index 5116a9b934fd8..6f87fb13988bf 100644 --- a/Mathlib/CategoryTheory/Yoneda.lean +++ b/Mathlib/CategoryTheory/Yoneda.lean @@ -582,6 +582,19 @@ lemma isIso_of_yoneda_map_bijective {X Y : C} (f : X ⟶ Y) obtain ⟨g, hg : g ≫ f = 𝟙 Y⟩ := (hf Y).2 (𝟙 Y) exact ⟨g, (hf _).1 (by aesop_cat), hg⟩ +lemma isIso_iff_yoneda_map_bijective {X Y : C} (f : X ⟶ Y) : + IsIso f ↔ (∀ (T : C), Function.Bijective (fun (x : T ⟶ X) => x ≫ f)) := by + refine ⟨fun _ ↦ ?_, fun hf ↦ isIso_of_yoneda_map_bijective f hf⟩ + have : IsIso (yoneda.map f) := inferInstance + intro T + rw [← isIso_iff_bijective] + exact inferInstanceAs (IsIso ((yoneda.map f).app _)) + +lemma isIso_iff_isIso_yoneda_map {X Y : C} (f : X ⟶ Y) : + IsIso f ↔ ∀ c : C, IsIso ((yoneda.map f).app ⟨c⟩) := by + rw [isIso_iff_yoneda_map_bijective] + exact forall_congr' fun _ ↦ (isIso_iff_bijective _).symm + end YonedaLemma section CoyonedaLemma @@ -749,6 +762,19 @@ lemma isIso_of_coyoneda_map_bijective {X Y : C} (f : X ⟶ Y) refine ⟨g, hg, (hf _).1 ?_⟩ simp only [Category.comp_id, ← Category.assoc, hg, Category.id_comp] +lemma isIso_iff_coyoneda_map_bijective {X Y : C} (f : X ⟶ Y) : + IsIso f ↔ (∀ (T : C), Function.Bijective (fun (x : Y ⟶ T) => f ≫ x)) := by + refine ⟨fun _ ↦ ?_, fun hf ↦ isIso_of_coyoneda_map_bijective f hf⟩ + have : IsIso (coyoneda.map f.op) := inferInstance + intro T + rw [← isIso_iff_bijective] + exact inferInstanceAs (IsIso ((coyoneda.map f.op).app _)) + +lemma isIso_iff_isIso_coyoneda_map {X Y : C} (f : X ⟶ Y) : + IsIso f ↔ ∀ c : C, IsIso ((coyoneda.map f.op).app c) := by + rw [isIso_iff_coyoneda_map_bijective] + exact forall_congr' fun _ ↦ (isIso_iff_bijective _).symm + end CoyonedaLemma section diff --git a/docs/references.bib b/docs/references.bib index ecdb8ed7428cd..f358fdab188cc 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -1038,6 +1038,20 @@ @Book{ davey_priestley url = {https://doi.org/10.1017/CBO9780511809088} } +@Article{ day1972, + author = {Day, Brian}, + title = {A reflection theorem for closed categories}, + journal = {J. Pure Appl. Algebra}, + fjournal = {Journal of Pure and Applied Algebra}, + volume = {2}, + year = {1972}, + number = {1}, + pages = {1--11}, + issn = {0022-4049}, + doi = {10.1016/0022-4049(72)90021-7}, + url = {https://doi.org/10.1016/0022-4049(72)90021-7} +} + @InProceedings{ deligne_formulaire, author = {Deligne, P.}, title = {Courbes elliptiques: formulaire d'apr\`es {J}. {T}ate}, From dd8f6e1179f68dfd7a939b270ede7385437e4836 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sun, 27 Oct 2024 10:17:44 +0000 Subject: [PATCH 471/505] feat: add `MultilinearMap.dfinsuppFamily` and `MultilinearMap.piFamily` (#17881) This is the analog to `Fintype.piFinset` Since this is computable and rather confusing, I've included a test with a concrete evaluation. Co-authored-by: Eric Wieser --- Mathlib.lean | 2 + .../LinearAlgebra/Multilinear/DFinsupp.lean | 144 ++++++++++++++++++ Mathlib/LinearAlgebra/Multilinear/Pi.lean | 112 ++++++++++++++ test/DFinsuppMultiLinear.lean | 21 +++ 4 files changed, 279 insertions(+) create mode 100644 Mathlib/LinearAlgebra/Multilinear/DFinsupp.lean create mode 100644 Mathlib/LinearAlgebra/Multilinear/Pi.lean create mode 100644 test/DFinsuppMultiLinear.lean diff --git a/Mathlib.lean b/Mathlib.lean index eb696d6f93a41..7484c32878cb3 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3189,7 +3189,9 @@ import Mathlib.LinearAlgebra.Matrix.Transvection import Mathlib.LinearAlgebra.Matrix.ZPow import Mathlib.LinearAlgebra.Multilinear.Basic import Mathlib.LinearAlgebra.Multilinear.Basis +import Mathlib.LinearAlgebra.Multilinear.DFinsupp import Mathlib.LinearAlgebra.Multilinear.FiniteDimensional +import Mathlib.LinearAlgebra.Multilinear.Pi import Mathlib.LinearAlgebra.Multilinear.TensorProduct import Mathlib.LinearAlgebra.Orientation import Mathlib.LinearAlgebra.PID diff --git a/Mathlib/LinearAlgebra/Multilinear/DFinsupp.lean b/Mathlib/LinearAlgebra/Multilinear/DFinsupp.lean new file mode 100644 index 0000000000000..b773dab0992f7 --- /dev/null +++ b/Mathlib/LinearAlgebra/Multilinear/DFinsupp.lean @@ -0,0 +1,144 @@ +/- +Copyright (c) 2024 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser, Sophie Morel +-/ +import Mathlib.LinearAlgebra.DFinsupp +import Mathlib.LinearAlgebra.Multilinear.Basic + +/-! +# Interactions between finitely-supported functions and multilinear maps + +This file provides `MultilinearMap.dfinsuppFamily`, which satisfies +`MultilinearMap.dfinsuppFamily f x p = f p (fun i => x i (p i))`. +This is the finitely-supported version of `MultilinearMap.piFamily`. + +This is useful because all the intermediate results are bundled: + +* `MultilinearMap.dfinsuppFamily f x` is a `DFinsupp` supported by families of indices `p`. +* `MultilinearMap.dfinsuppFamily f` is a `MultilinearMap` operating on finitely-supported functions + `x`. +* `MultilinearMap.dfinsuppFamilyₗ` is a `LinearMap`, linear in the family of multilinear maps `f`. + +-/ + +universe uι uκ uS uR uM uN +variable {ι : Type uι} {κ : ι → Type uκ} +variable {S : Type uS} {R : Type uR} {M : ∀ i, κ i → Type uM} {N : (Π i, κ i) → Type uN} + +namespace MultilinearMap + +section Semiring + +variable [DecidableEq ι] [Fintype ι] [Semiring R] +variable [∀ i k, AddCommMonoid (M i k)] [∀ p, AddCommMonoid (N p)] +variable [∀ i k, Module R (M i k)] [∀ p, Module R (N p)] + +/-- +Given a family of indices `κ` and a multilinear map `f p` for each way `p` to select one index from +each family, `dfinsuppFamily f` maps a family of finitely-supported functions (one for each domain +`κ i`) into a finitely-supported function from each selection of indices (with domain `Π i, κ i`). + +Strictly this doesn't need multilinearity, only the fact that `f p m = 0` whenever `m i = 0` for +some `i`. + +This is the `DFinsupp` version of `MultilinearMap.pi'`. +-/ +@[simps] +def dfinsuppFamily + (f : Π (p : Π i, κ i), MultilinearMap R (fun i ↦ M i (p i)) (N p)) : + MultilinearMap R (fun i => Π₀ j : κ i, M i j) (Π₀ t : Π i, κ i, N t) where + toFun x := + { toFun := fun p => f p (fun i => x i (p i)) + support' := (Trunc.finChoice fun i => (x i).support').map fun s => ⟨ + Finset.univ.val.pi (fun i ↦ (s i).val) |>.map fun f i => f i (Finset.mem_univ _), + fun p => by + simp only [Multiset.mem_map, Multiset.mem_pi, Finset.mem_val, Finset.mem_univ, + forall_true_left] + simp_rw [or_iff_not_imp_right] + intro h + push_neg at h + refine ⟨fun i _ => p i, fun i => (s i).prop _ |>.resolve_right ?_, rfl⟩ + exact mt ((f p).map_coord_zero (m := fun i => x i _) i) h⟩} + map_add' {dec} m i x y := DFinsupp.ext fun p => by + cases Subsingleton.elim dec (by infer_instance) + dsimp + simp_rw [Function.apply_update (fun i m => m (p i)) m, DFinsupp.add_apply, (f p).map_add] + map_smul' {dec} m i c x := DFinsupp.ext fun p => by + cases Subsingleton.elim dec (by infer_instance) + dsimp + simp_rw [Function.apply_update (fun i m => m (p i)) m, DFinsupp.smul_apply, (f p).map_smul] + +theorem support_dfinsuppFamily_subset + [∀ i, DecidableEq (κ i)] + [∀ i j, (x : M i j) → Decidable (x ≠ 0)] [∀ i, (x : N i) → Decidable (x ≠ 0)] + (f : Π (p : Π i, κ i), MultilinearMap R (fun i ↦ M i (p i)) (N p)) + (x : ∀ i, Π₀ j : κ i, M i j) : + (dfinsuppFamily f x).support ⊆ Fintype.piFinset fun i => (x i).support := by + intro p hp + simp only [DFinsupp.mem_support_toFun, dfinsuppFamily_apply_toFun, ne_eq, + Fintype.mem_piFinset] at hp ⊢ + intro i + exact mt ((f p).map_coord_zero (m := fun i => x i _) i) hp + +/-- When applied to a family of finitely-supported functions each supported on a single element, +`dfinsuppFamily` is itself supported on a single element, with value equal to the map `f` applied +at that point. -/ +@[simp] +theorem dfinsuppFamily_single [∀ i, DecidableEq (κ i)] + (f : Π (p : Π i, κ i), MultilinearMap R (fun i ↦ M i (p i)) (N p)) + (p : ∀ i, κ i) (m : ∀ i, M i (p i)) : + dfinsuppFamily f (fun i => .single (p i) (m i)) = DFinsupp.single p (f p m) := by + ext q + obtain rfl | hpq := eq_or_ne p q + · simp + · rw [DFinsupp.single_eq_of_ne hpq] + rw [Function.ne_iff] at hpq + obtain ⟨i, hpqi⟩ := hpq + apply (f q).map_coord_zero i + simp_rw [DFinsupp.single_eq_of_ne hpqi] + +@[simp] +theorem dfinsuppFamily_compLinearMap_lsingle [∀ i, DecidableEq (κ i)] + (f : Π (p : Π i, κ i), MultilinearMap R (fun i ↦ M i (p i)) (N p)) (p : ∀ i, κ i) : + (dfinsuppFamily f).compLinearMap (fun i => DFinsupp.lsingle (p i)) + = (DFinsupp.lsingle p).compMultilinearMap (f p) := + MultilinearMap.ext <| dfinsuppFamily_single f p + +@[simp] +theorem dfinsuppFamily_zero : + dfinsuppFamily (0 : Π (p : Π i, κ i), MultilinearMap R (fun i ↦ M i (p i)) (N p)) = 0 := by + ext; simp + +@[simp] +theorem dfinsuppFamily_add (f g : Π (p : Π i, κ i), MultilinearMap R (fun i ↦ M i (p i)) (N p)) : + dfinsuppFamily (f + g) = dfinsuppFamily f + dfinsuppFamily g := by + ext; simp + +@[simp] +theorem dfinsuppFamily_smul + [Monoid S] [∀ p, DistribMulAction S (N p)] [∀ p, SMulCommClass R S (N p)] + (s : S) (f : Π (p : Π i, κ i), MultilinearMap R (fun i ↦ M i (p i)) (N p)) : + dfinsuppFamily (s • f) = s • dfinsuppFamily f := by + ext; simp + +end Semiring + +section CommSemiring + +variable [DecidableEq ι] [Fintype ι] [CommSemiring R] +variable [∀ i k, AddCommMonoid (M i k)] [∀ p, AddCommMonoid (N p)] +variable [∀ i k, Module R (M i k)] [∀ p, Module R (N p)] + +/-- `MultilinearMap.dfinsuppFamily` as a linear map. -/ +@[simps] +def dfinsuppFamilyₗ : + (Π (p : Π i, κ i), MultilinearMap R (fun i ↦ M i (p i)) (N p)) + →ₗ[R] MultilinearMap R (fun i => Π₀ j : κ i, M i j) (Π₀ t : Π i, κ i, N t) where + toFun := dfinsuppFamily + map_add' := dfinsuppFamily_add + map_smul' := dfinsuppFamily_smul + +end CommSemiring + +end MultilinearMap diff --git a/Mathlib/LinearAlgebra/Multilinear/Pi.lean b/Mathlib/LinearAlgebra/Multilinear/Pi.lean new file mode 100644 index 0000000000000..09760ac3cb1bb --- /dev/null +++ b/Mathlib/LinearAlgebra/Multilinear/Pi.lean @@ -0,0 +1,112 @@ +/- +Copyright (c) 2024 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser +-/ +import Mathlib.LinearAlgebra.Pi +import Mathlib.LinearAlgebra.Multilinear.Basic + +/-! +# Interactions between (dependent) functions and multilinear maps + +This file provides `MultilinearMap.pi`, which satisfies +`piFamily f x p = f p (fun i => x i (p i))`. + +This is useful because all the intermediate results are bundled: + +* `MultilinearMap.pi f` is a `MultilinearMap` operating on functions `x`. +* `MultilinearMap.piₗ` is a `LinearMap`, linear in the family of multilinear maps `f`. + +-/ + +universe uι uκ uS uR uM uN +variable {ι : Type uι} {κ : ι → Type uκ} +variable {S : Type uS} {R : Type uR} {M : ∀ i, κ i → Type uM} {N : (Π i, κ i) → Type uN} + +namespace MultilinearMap + +section Semiring + +variable [Semiring R] +variable [∀ i k, AddCommMonoid (M i k)] [∀ p, AddCommMonoid (N p)] +variable [∀ i k, Module R (M i k)] [∀ p, Module R (N p)] + +/-- +Given a family of indices `κ` and a multilinear map `f p` for each way `p` to select one index from +each family, `piFamily f` maps a family of functions (one for each domain `κ i`) into a function +from each selection of indices (with domain `Π i, κ i`). +-/ +@[simps] +def piFamily (f : Π (p : Π i, κ i), MultilinearMap R (fun i ↦ M i (p i)) (N p)) : + MultilinearMap R (fun i => Π j : κ i, M i j) (Π t : Π i, κ i, N t) where + toFun x := fun p => f p (fun i => x i (p i)) + map_add' {dec} m i x y := funext fun p => by + cases Subsingleton.elim dec (by infer_instance) + dsimp + simp_rw [Function.apply_update (fun i m => m (p i)) m, Pi.add_apply, (f p).map_add] + map_smul' {dec} m i c x := funext fun p => by + cases Subsingleton.elim dec (by infer_instance) + dsimp + simp_rw [Function.apply_update (fun i m => m (p i)) m, Pi.smul_apply, (f p).map_smul] + +/-- When applied to a family of finitely-supported functions each supported on a single element, +`piFamily` is itself supported on a single element, with value equal to the map `f` applied +at that point. -/ +@[simp] +theorem piFamily_single [Fintype ι] [∀ i, DecidableEq (κ i)] + (f : Π (p : Π i, κ i), MultilinearMap R (fun i ↦ M i (p i)) (N p)) + (p : ∀ i, κ i) (m : ∀ i, M i (p i)) : + piFamily f (fun i => Pi.single (p i) (m i)) = Pi.single p (f p m) := by + ext q + obtain rfl | hpq := eq_or_ne p q + · simp + · rw [Pi.single_eq_of_ne' hpq] + rw [Function.ne_iff] at hpq + obtain ⟨i, hpqi⟩ := hpq + apply (f q).map_coord_zero i + simp_rw [Pi.single_eq_of_ne' hpqi] + +@[simp] +theorem piFamily_compLinearMap_lsingle [Fintype ι] [∀ i, DecidableEq (κ i)] + (f : Π (p : Π i, κ i), MultilinearMap R (fun i ↦ M i (p i)) (N p)) (p : ∀ i, κ i) : + (piFamily f).compLinearMap (fun i => LinearMap.single _ _ (p i)) + = (LinearMap.single _ _ p).compMultilinearMap (f p) := + MultilinearMap.ext <| piFamily_single f p + +@[simp] +theorem piFamily_zero : + piFamily (0 : Π (p : Π i, κ i), MultilinearMap R (fun i ↦ M i (p i)) (N p)) = 0 := by + ext; simp + +@[simp] +theorem piFamily_add (f g : Π (p : Π i, κ i), MultilinearMap R (fun i ↦ M i (p i)) (N p)) : + piFamily (f + g) = piFamily f + piFamily g := by + ext; simp + +@[simp] +theorem piFamily_smul + [Monoid S] [∀ p, DistribMulAction S (N p)] [∀ p, SMulCommClass R S (N p)] + (s : S) (f : Π (p : Π i, κ i), MultilinearMap R (fun i ↦ M i (p i)) (N p)) : + piFamily (s • f) = s • piFamily f := by + ext; simp + +end Semiring + +section CommSemiring + +variable [CommSemiring R] +variable [∀ i k, AddCommMonoid (M i k)] [∀ p, AddCommMonoid (N p)] +variable [∀ i k, Module R (M i k)] [∀ p, Module R (N p)] + +/-- `MultilinearMap.piFamily` as a linear map. -/ +@[simps] +def piFamilyₗ : + (Π (p : Π i, κ i), MultilinearMap R (fun i ↦ M i (p i)) (N p)) + →ₗ[R] MultilinearMap R (fun i => Π j : κ i, M i j) (Π t : Π i, κ i, N t) where + toFun := piFamily + map_add' := piFamily_add + map_smul' := piFamily_smul + +end CommSemiring + +end MultilinearMap diff --git a/test/DFinsuppMultiLinear.lean b/test/DFinsuppMultiLinear.lean new file mode 100644 index 0000000000000..41ccc8e876c46 --- /dev/null +++ b/test/DFinsuppMultiLinear.lean @@ -0,0 +1,21 @@ +import Mathlib.LinearAlgebra.Multilinear.DFinsupp +import Mathlib.LinearAlgebra.Multilinear.Pi +import Mathlib.Data.DFinsupp.Notation + +/-- +info: fun₀ + | !["hello", "complicated", "world"] => 20 + | !["hello", "complicated", "test file"] => 30 + | !["goodbye", "complicated", "world"] => -20 + | !["goodbye", "complicated", "test file"] => -30 +-/ +#guard_msgs in +#eval MultilinearMap.dfinsuppFamily + (κ := fun _ => String) + (M := fun _ _ => ℤ) + (N := fun _ => ℤ) + (R := ℤ) + (f := fun _ => MultilinearMap.mkPiAlgebra ℤ (Fin 3) ℤ) ![ + fun₀ | "hello" => 1 | "goodbye" => -1, + fun₀ | "complicated" => 10, + fun₀ | "world" => 2 | "test file" => 3] From e02ca3c69dcad5509a898ba89b1cc9e72d74b777 Mon Sep 17 00:00:00 2001 From: "Tristan F.-R." Date: Sun, 27 Oct 2024 10:17:45 +0000 Subject: [PATCH 472/505] doc(Data/List/Permutation): remove outdated TODO (#18274) Covered in https://github.com/leanprover-community/mathlib4/pull/15512. Co-authored-by: LeoDog896 --- Mathlib/Data/List/Permutation.lean | 4 ---- 1 file changed, 4 deletions(-) diff --git a/Mathlib/Data/List/Permutation.lean b/Mathlib/Data/List/Permutation.lean index 2798655fad89c..7b30383ce32a7 100644 --- a/Mathlib/Data/List/Permutation.lean +++ b/Mathlib/Data/List/Permutation.lean @@ -42,10 +42,6 @@ all positions. Hence, to build `[0, 1, 2, 3].permutations'`, it does `[0, 1, 3, 2], [1, 0, 3, 2], [1, 3, 0, 2], [1, 3, 2, 0],` `[0, 3, 1, 2], [3, 0, 1, 2], [3, 1, 0, 2], [3, 1, 2, 0],` `[0, 3, 2, 1], [3, 0, 2, 1], [3, 2, 0, 1], [3, 2, 1, 0]]` - -## TODO - -Show that `l.Nodup → l.permutations.Nodup`. See `Data.Fintype.List`. -/ -- Make sure we don't import algebra From 1eb2db95be7a6eb668f2a1c82d1a4f4c7111db7c Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sun, 27 Oct 2024 10:54:19 +0000 Subject: [PATCH 473/505] fix(Algebra/MonoidAlgebra/Basic): resolve `ext` porting notes (#6524) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Yaël Dillies --- Mathlib/Algebra/MonoidAlgebra/Basic.lean | 15 +- Mathlib/Algebra/MonoidAlgebra/Defs.lean | 167 +++++++++++++++++++---- Mathlib/Algebra/MvPolynomial/Basic.lean | 2 +- Mathlib/Data/Finsupp/Basic.lean | 4 +- 4 files changed, 145 insertions(+), 43 deletions(-) diff --git a/Mathlib/Algebra/MonoidAlgebra/Basic.lean b/Mathlib/Algebra/MonoidAlgebra/Basic.lean index 9eb40ad475be0..e5d8512805835 100644 --- a/Mathlib/Algebra/MonoidAlgebra/Basic.lean +++ b/Mathlib/Algebra/MonoidAlgebra/Basic.lean @@ -90,8 +90,7 @@ def liftMagma [Module k A] [IsScalarTower k A A] [SMulCommClass k A A] : sum_single_index, Function.comp_apply, one_smul, zero_smul, MulHom.coe_comp, NonUnitalAlgHom.coe_to_mulHom] right_inv F := by - -- Porting note (#11041): `ext` → `refine nonUnitalAlgHom_ext' k (MulHom.ext fun m => ?_)` - refine nonUnitalAlgHom_ext' k (MulHom.ext fun m => ?_) + ext m simp only [NonUnitalAlgHom.coe_mk, ofMagma_apply, NonUnitalAlgHom.toMulHom_eq_coe, sum_single_index, Function.comp_apply, one_smul, zero_smul, MulHom.coe_comp, NonUnitalAlgHom.coe_to_mulHom] @@ -109,9 +108,8 @@ In particular this provides the instance `Algebra k (MonoidAlgebra k G)`. instance algebra {A : Type*} [CommSemiring k] [Semiring A] [Algebra k A] [Monoid G] : Algebra k (MonoidAlgebra A G) := { singleOneRingHom.comp (algebraMap k A) with - -- Porting note (#11041): `ext` → `refine Finsupp.ext fun _ => ?_` smul_def' := fun r a => by - refine Finsupp.ext fun _ => ?_ + ext -- Porting note: Newly required. rw [Finsupp.coe_smul] simp [single_one_mul_apply, Algebra.smul_def, Pi.smul_apply] @@ -125,8 +123,7 @@ def singleOneAlgHom {A : Type*} [CommSemiring k] [Semiring A] [Algebra k A] [Mon A →ₐ[k] MonoidAlgebra A G := { singleOneRingHom with commutes' := fun r => by - -- Porting note (#11041): `ext` → `refine Finsupp.ext fun _ => ?_` - refine Finsupp.ext fun _ => ?_ + ext simp rfl } @@ -390,9 +387,8 @@ In particular this provides the instance `Algebra k k[G]`. instance algebra [CommSemiring R] [Semiring k] [Algebra R k] [AddMonoid G] : Algebra R k[G] := { singleZeroRingHom.comp (algebraMap R k) with - -- Porting note (#11041): `ext` → `refine Finsupp.ext fun _ => ?_` smul_def' := fun r a => by - refine Finsupp.ext fun _ => ?_ + ext -- Porting note: Newly required. rw [Finsupp.coe_smul] simp [single_zero_mul_apply, Algebra.smul_def, Pi.smul_apply] @@ -405,8 +401,7 @@ instance algebra [CommSemiring R] [Semiring k] [Algebra R k] [AddMonoid G] : def singleZeroAlgHom [CommSemiring R] [Semiring k] [Algebra R k] [AddMonoid G] : k →ₐ[R] k[G] := { singleZeroRingHom with commutes' := fun r => by - -- Porting note (#11041): `ext` → `refine Finsupp.ext fun _ => ?_` - refine Finsupp.ext fun _ => ?_ + ext simp rfl } diff --git a/Mathlib/Algebra/MonoidAlgebra/Defs.lean b/Mathlib/Algebra/MonoidAlgebra/Defs.lean index 3766874adfc29..4da403f3dc92a 100644 --- a/Mathlib/Algebra/MonoidAlgebra/Defs.lean +++ b/Mathlib/Algebra/MonoidAlgebra/Defs.lean @@ -377,6 +377,63 @@ def comapDistribMulActionSelf [Group G] [Semiring k] : DistribMulAction G (Monoi end DerivedInstances +/-! +#### Copies of `ext` lemmas and bundled `single`s from `Finsupp` + +As `MonoidAlgebra` is a type synonym, `ext` will not unfold it to find `ext` lemmas. +We need bundled version of `Finsupp.single` with the right types to state these lemmas. +It is good practice to have those, regardless of the `ext` issue. +-/ + +section ExtLemmas + +/-- A copy of `Finsupp.ext` for `MonoidAlgebra`. -/ +@[ext] +theorem ext [Semiring k] ⦃f g : MonoidAlgebra k G⦄ (H : ∀ (x : G), f x = g x) : + f = g := + Finsupp.ext H + +/-- A copy of `Finsupp.singleAddHom` for `MonoidAlgebra`. -/ +abbrev singleAddHom [Semiring k] (a : G) : k →+ MonoidAlgebra k G := Finsupp.singleAddHom a + +@[simp] lemma singleAddHom_apply [Semiring k] (a : G) (b : k) : + singleAddHom a b = single a b := rfl + +/-- A copy of `Finsupp.addHom_ext'` for `MonoidAlgebra`. -/ +@[ext high] +theorem addHom_ext' {N : Type*} [Semiring k] [AddZeroClass N] + ⦃f g : MonoidAlgebra k G →+ N⦄ + (H : ∀ (x : G), AddMonoidHom.comp f (singleAddHom x) = AddMonoidHom.comp g (singleAddHom x)) : + f = g := + Finsupp.addHom_ext' H + +/-- A copy of `Finsupp.distribMulActionHom_ext'` for `MonoidAlgebra`. -/ +@[ext] +theorem distribMulActionHom_ext' {N : Type*} [Monoid R] [Semiring k] [AddMonoid N] + [DistribMulAction R N] [DistribMulAction R k] + {f g : MonoidAlgebra k G →+[R] N} + (h : ∀ a : G, + f.comp (DistribMulActionHom.single (M := k) a) = g.comp (DistribMulActionHom.single a)) : + f = g := + Finsupp.distribMulActionHom_ext' h + +/-- A copy of `Finsupp.lsingle` for `MonoidAlgebra`. -/ +abbrev lsingle [Semiring R] [Semiring k] [Module R k] (a : G) : + k →ₗ[R] MonoidAlgebra k G := Finsupp.lsingle a + +@[simp] lemma lsingle_apply [Semiring R] [Semiring k] [Module R k] (a : G) (b : k) : + lsingle (R := R) a b = single a b := rfl + +/-- A copy of `Finsupp.lhom_ext'` for `MonoidAlgebra`. -/ +@[ext high] +lemma lhom_ext' {N : Type*} [Semiring R] [Semiring k] [AddCommMonoid N] [Module R N] [Module R k] + ⦃f g : MonoidAlgebra k G →ₗ[R] N⦄ + (H : ∀ (x : G), LinearMap.comp f (lsingle x) = LinearMap.comp g (lsingle x)) : + f = g := + Finsupp.lhom_ext' H + +end ExtLemmas + section MiscTheorems variable [Semiring k] @@ -563,12 +620,11 @@ theorem liftNC_smul [MulOneClass G] {R : Type*} [Semiring R] (f : k →+* R) (g suffices (liftNC (↑f) g).comp (smulAddHom k (MonoidAlgebra k G) c) = (AddMonoidHom.mulLeft (f c)).comp (liftNC (↑f) g) from DFunLike.congr_fun this φ - -- Porting note (#11041): `ext` couldn't a find appropriate theorem. - refine addHom_ext' fun a => AddMonoidHom.ext fun b => ?_ + ext -- Porting note: `reducible` cannot be `local` so the proof gets more complex. unfold MonoidAlgebra simp only [AddMonoidHom.coe_comp, Function.comp_apply, singleAddHom_apply, smulAddHom_apply, - smul_single, smul_eq_mul, AddMonoidHom.coe_mulLeft] + smul_single, smul_eq_mul, AddMonoidHom.coe_mulLeft, Finsupp.singleAddHom_apply] -- This used to be `rw`, but we need `erw` after leanprover/lean4#2644 erw [liftNC_single, liftNC_single]; rw [AddMonoidHom.coe_coe, map_mul, mul_assoc] @@ -582,34 +638,32 @@ section NonUnitalNonAssocAlgebra variable (k) [Semiring k] [DistribSMul R k] [Mul G] instance isScalarTower_self [IsScalarTower R k k] : - IsScalarTower R (MonoidAlgebra k G) (MonoidAlgebra k G) := - ⟨fun t a b => by - -- Porting note (#11041): `ext` → `refine Finsupp.ext fun _ => ?_` - refine Finsupp.ext fun m => ?_ + IsScalarTower R (MonoidAlgebra k G) (MonoidAlgebra k G) where + smul_assoc t a b := by + ext -- Porting note: `refine` & `rw` are required because `simp` behaves differently. classical - simp only [smul_eq_mul, mul_apply] - rw [coe_smul] - refine Eq.trans (sum_smul_index' (g := a) (b := t) ?_) ?_ <;> - simp only [mul_apply, Finsupp.smul_sum, smul_ite, smul_mul_assoc, - zero_mul, ite_self, imp_true_iff, sum_zero, Pi.smul_apply, smul_zero]⟩ + simp only [smul_eq_mul, mul_apply] + rw [coe_smul] + refine Eq.trans (sum_smul_index' (g := a) (b := t) ?_) ?_ <;> + simp only [mul_apply, Finsupp.smul_sum, smul_ite, smul_mul_assoc, + zero_mul, ite_self, imp_true_iff, sum_zero, Pi.smul_apply, smul_zero] /-- Note that if `k` is a `CommSemiring` then we have `SMulCommClass k k k` and so we can take `R = k` in the below. In other words, if the coefficients are commutative amongst themselves, they also commute with the algebra multiplication. -/ instance smulCommClass_self [SMulCommClass R k k] : - SMulCommClass R (MonoidAlgebra k G) (MonoidAlgebra k G) := - ⟨fun t a b => by - -- Porting note (#11041): `ext` → `refine Finsupp.ext fun _ => ?_` - refine Finsupp.ext fun m => ?_ + SMulCommClass R (MonoidAlgebra k G) (MonoidAlgebra k G) where + smul_comm t a b := by + ext -- Porting note: `refine` & `rw` are required because `simp` behaves differently. classical - simp only [smul_eq_mul, mul_apply] - rw [coe_smul] - refine Eq.symm (Eq.trans (congr_arg (sum a) - (funext₂ fun a₁ b₁ => sum_smul_index' (g := b) (b := t) ?_)) ?_) <;> - simp only [mul_apply, Finsupp.sum, Finset.smul_sum, smul_ite, mul_smul_comm, - imp_true_iff, ite_eq_right_iff, Pi.smul_apply, mul_zero, smul_zero]⟩ + simp only [smul_eq_mul, mul_apply] + rw [coe_smul] + refine Eq.symm (Eq.trans (congr_arg (sum a) + (funext₂ fun a₁ b₁ => sum_smul_index' (g := b) (b := t) ?_)) ?_) <;> + simp only [mul_apply, Finsupp.sum, Finset.smul_sum, smul_ite, mul_smul_comm, + imp_true_iff, ite_eq_right_iff, Pi.smul_apply, mul_zero, smul_zero] instance smulCommClass_symm_self [SMulCommClass k R k] : SMulCommClass (MonoidAlgebra k G) R (MonoidAlgebra k G) := @@ -628,9 +682,7 @@ theorem single_one_comm [CommSemiring k] [MulOneClass G] (r : k) (f : MonoidAlge def singleOneRingHom [Semiring k] [MulOneClass G] : k →+* MonoidAlgebra k G := { Finsupp.singleAddHom 1 with map_one' := rfl - map_mul' := fun x y => by - simp only [ZeroHom.toFun_eq_coe, AddMonoidHom.toZeroHom_coe, singleAddHom_apply, - single_mul_single, mul_one] } + map_mul' := fun x y => by simp } /-- If `f : G → H` is a multiplicative homomorphism between two monoids, then `Finsupp.mapDomain f` is a ring homomorphism between their monoid algebras. -/ @@ -740,9 +792,7 @@ protected noncomputable def opRingEquiv [Monoid G] : rw [← AddEquiv.coe_toAddMonoidHom] refine Iff.mpr (AddMonoidHom.map_mul_iff (R := (MonoidAlgebra k G)ᵐᵒᵖ) (S := MonoidAlgebra kᵐᵒᵖ Gᵐᵒᵖ) _) ?_ - -- Porting note (#11041): Was `ext`. - refine AddMonoidHom.mul_op_ext _ _ <| addHom_ext' fun i₁ => AddMonoidHom.ext fun r₁ => - AddMonoidHom.mul_op_ext _ _ <| addHom_ext' fun i₂ => AddMonoidHom.ext fun r₂ => ?_ + ext -- Porting note: `reducible` cannot be `local` so proof gets long. simp only [AddMonoidHom.coe_comp, AddEquiv.coe_toAddMonoidHom, opAddEquiv_apply, Function.comp_apply, singleAddHom_apply, AddMonoidHom.compr₂_apply, AddMonoidHom.coe_mul, @@ -1118,6 +1168,63 @@ because we've never discussed actions of additive groups. -/ end DerivedInstances +/-! +#### Copies of `ext` lemmas and bundled `single`s from `Finsupp` + +As `AddMonoidAlgebra` is a type synonym, `ext` will not unfold it to find `ext` lemmas. +We need bundled version of `Finsupp.single` with the right types to state these lemmas. +It is good practice to have those, regardless of the `ext` issue. +-/ + +section ExtLemmas + +/-- A copy of `Finsupp.ext` for `AddMonoidAlgebra`. -/ +@[ext] +theorem ext [Semiring k] ⦃f g : AddMonoidAlgebra k G⦄ (H : ∀ (x : G), f x = g x) : + f = g := + Finsupp.ext H + +/-- A copy of `Finsupp.singleAddHom` for `AddMonoidAlgebra`. -/ +abbrev singleAddHom [Semiring k] (a : G) : k →+ AddMonoidAlgebra k G := Finsupp.singleAddHom a + +@[simp] lemma singleAddHom_apply [Semiring k] (a : G) (b : k) : + singleAddHom a b = single a b := rfl + +/-- A copy of `Finsupp.addHom_ext'` for `AddMonoidAlgebra`. -/ +@[ext high] +theorem addHom_ext' {N : Type*} [Semiring k] [AddZeroClass N] + ⦃f g : AddMonoidAlgebra k G →+ N⦄ + (H : ∀ (x : G), AddMonoidHom.comp f (singleAddHom x) = AddMonoidHom.comp g (singleAddHom x)) : + f = g := + Finsupp.addHom_ext' H + +/-- A copy of `Finsupp.distribMulActionHom_ext'` for `AddMonoidAlgebra`. -/ +@[ext] +theorem distribMulActionHom_ext' {N : Type*} [Monoid R] [Semiring k] [AddMonoid N] + [DistribMulAction R N] [DistribMulAction R k] + {f g : AddMonoidAlgebra k G →+[R] N} + (h : ∀ a : G, + f.comp (DistribMulActionHom.single (M := k) a) = g.comp (DistribMulActionHom.single a)) : + f = g := + Finsupp.distribMulActionHom_ext' h + +/-- A copy of `Finsupp.lsingle` for `AddMonoidAlgebra`. -/ +abbrev lsingle [Semiring R] [Semiring k] [Module R k] (a : G) : + k →ₗ[R] AddMonoidAlgebra k G := Finsupp.lsingle a + +@[simp] lemma lsingle_apply [Semiring R] [Semiring k] [Module R k] (a : G) (b : k) : + lsingle (R := R) a b = single a b := rfl + +/-- A copy of `Finsupp.lhom_ext'` for `AddMonoidAlgebra`. -/ +@[ext high] +lemma lhom_ext' {N : Type*} [Semiring R] [Semiring k] [AddCommMonoid N] [Module R N] [Module R k] + ⦃f g : AddMonoidAlgebra k G →ₗ[R] N⦄ + (H : ∀ (x : G), LinearMap.comp f (lsingle x) = LinearMap.comp g (lsingle x)) : + f = g := + Finsupp.lhom_ext' H + +end ExtLemmas + section MiscTheorems variable [Semiring k] @@ -1363,7 +1470,7 @@ section Algebra def singleZeroRingHom [Semiring k] [AddMonoid G] : k →+* k[G] := { Finsupp.singleAddHom 0 with map_one' := rfl - map_mul' := fun x y => by simp only [singleAddHom, single_mul_single, zero_add] } + map_mul' := fun x y => by simp only [Finsupp.singleAddHom, single_mul_single, zero_add] } /-- If two ring homomorphisms from `k[G]` are equal on all `single a 1` and `single 0 b`, then they are equal. -/ @@ -1443,3 +1550,5 @@ theorem prod_single [CommSemiring k] [AddCommMonoid G] {s : Finset ι} {a : ι end end AddMonoidAlgebra + +set_option linter.style.longFile 1700 diff --git a/Mathlib/Algebra/MvPolynomial/Basic.lean b/Mathlib/Algebra/MvPolynomial/Basic.lean index 1e8efa3f5e687..93c6613dfbc95 100644 --- a/Mathlib/Algebra/MvPolynomial/Basic.lean +++ b/Mathlib/Algebra/MvPolynomial/Basic.lean @@ -155,7 +155,7 @@ variable [CommSemiring R] [CommSemiring S₁] {p q : MvPolynomial σ R} /-- `monomial s a` is the monomial with coefficient `a` and exponents given by `s` -/ def monomial (s : σ →₀ ℕ) : R →ₗ[R] MvPolynomial σ R := - lsingle s + AddMonoidAlgebra.lsingle s theorem single_eq_monomial (s : σ →₀ ℕ) (a : R) : Finsupp.single s a = monomial s a := rfl diff --git a/Mathlib/Data/Finsupp/Basic.lean b/Mathlib/Data/Finsupp/Basic.lean index e2a9c7d382a55..97af1005ed711 100644 --- a/Mathlib/Data/Finsupp/Basic.lean +++ b/Mathlib/Data/Finsupp/Basic.lean @@ -1434,9 +1434,7 @@ instance noZeroSMulDivisors [Semiring R] [AddCommMonoid M] [Module R M] {ι : Ty Finsupp.ext fun i => (smul_eq_zero.mp (DFunLike.ext_iff.mp h i)).resolve_left hc⟩ section DistribMulActionSemiHom - -variable [Semiring R] -variable [AddCommMonoid M] [AddCommMonoid N] [DistribMulAction R M] [DistribMulAction R N] +variable [Monoid R] [AddMonoid M] [AddMonoid N] [DistribMulAction R M] [DistribMulAction R N] /-- `Finsupp.single` as a `DistribMulActionSemiHom`. From 1a0cb23be5d4c6eae78ac6ac9ae8f365b1b1b2d6 Mon Sep 17 00:00:00 2001 From: Jz Pan Date: Sun, 27 Oct 2024 12:33:38 +0000 Subject: [PATCH 474/505] chore(RingTheory/PowerSeries/Basic): rename `Polynomial.ToPowerSeries` -> `Polynomial.toPowerSeries` (#18265) According to naming convention, it should be `Polynomial.toPowerSeries`. It's already `MvPolynomial.toMvPowerSeries` for `MvPolynomial`. --- Mathlib/RingTheory/PowerSeries/Basic.lean | 6 ++++-- 1 file changed, 4 insertions(+), 2 deletions(-) diff --git a/Mathlib/RingTheory/PowerSeries/Basic.lean b/Mathlib/RingTheory/PowerSeries/Basic.lean index 8965ee6ddf242..795804f9c1eed 100644 --- a/Mathlib/RingTheory/PowerSeries/Basic.lean +++ b/Mathlib/RingTheory/PowerSeries/Basic.lean @@ -763,12 +763,14 @@ variable {R : Type*} [CommSemiring R] (φ ψ : R[X]) -- Porting note: added so we can add the `@[coe]` attribute /-- The natural inclusion from polynomials into formal power series. -/ @[coe] -def ToPowerSeries : R[X] → (PowerSeries R) := fun φ => +def toPowerSeries : R[X] → (PowerSeries R) := fun φ => PowerSeries.mk fun n => coeff φ n +@[deprecated (since := "2024-10-27")] alias ToPowerSeries := toPowerSeries + /-- The natural inclusion from polynomials into formal power series. -/ instance coeToPowerSeries : Coe R[X] (PowerSeries R) := - ⟨ToPowerSeries⟩ + ⟨toPowerSeries⟩ theorem coe_def : (φ : PowerSeries R) = PowerSeries.mk (coeff φ) := rfl From 6264905204f3ed595bc99032ec15af1910a0eddd Mon Sep 17 00:00:00 2001 From: Yakov Pechersky Date: Sun, 27 Oct 2024 13:11:53 +0000 Subject: [PATCH 475/505] feat(RingTheory/Lasker): decomposition into primary ideals (#17123) --- Mathlib.lean | 1 + Mathlib/RingTheory/Lasker.lean | 94 ++++++++++++++++++++++++++++++++++ 2 files changed, 95 insertions(+) create mode 100644 Mathlib/RingTheory/Lasker.lean diff --git a/Mathlib.lean b/Mathlib.lean index 7484c32878cb3..46d19d8a3517b 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4084,6 +4084,7 @@ import Mathlib.RingTheory.Kaehler.Polynomial import Mathlib.RingTheory.Kaehler.TensorProduct import Mathlib.RingTheory.KrullDimension.Basic import Mathlib.RingTheory.KrullDimension.Field +import Mathlib.RingTheory.Lasker import Mathlib.RingTheory.LaurentSeries import Mathlib.RingTheory.LittleWedderburn import Mathlib.RingTheory.LocalProperties.Basic diff --git a/Mathlib/RingTheory/Lasker.lean b/Mathlib/RingTheory/Lasker.lean new file mode 100644 index 0000000000000..f4f28d26fcd59 --- /dev/null +++ b/Mathlib/RingTheory/Lasker.lean @@ -0,0 +1,94 @@ +/- +Copyright (c) 2024 Yakov Pechersky. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yakov Pechersky +-/ +import Mathlib.Order.Irreducible +import Mathlib.RingTheory.Ideal.Colon +import Mathlib.RingTheory.Ideal.IsPrimary +import Mathlib.RingTheory.Noetherian + +/-! +# Lasker ring + +## Main declarations + +- `IsLasker`: A ring `R` satisfies `IsLasker R` when any `I : Ideal R` can be decomposed into + finitely many primary ideals. +- `Ideal.isLasker`: Every Noetherian commutative ring is a Lasker ring. + +## Implementation details + +There is a generalization for submodules that needs to be implemented. +Also, one needs to prove that the decomposition can be made minimal, +and that minimal decompositions are equivalent. + +-/ + +section IsLasker + +variable (R : Type*) [CommSemiring R] + +/-- A ring `R` satisfies `IsLasker R` when any `I : Ideal R` can be decomposed into +finitely many primary ideals.-/ +def IsLasker : Prop := + ∀ I : Ideal R, ∃ s : Finset (Ideal R), s.inf id = I ∧ ∀ ⦃J⦄, J ∈ s → J.IsPrimary + +end IsLasker + +namespace Ideal + +section Noetherian + +variable {R : Type*} [CommRing R] [IsNoetherianRing R] + +lemma _root_.InfIrred.isPrimary {I : Ideal R} (h : InfIrred I) : I.IsPrimary := by + refine ⟨h.ne_top, fun {a b} hab ↦ ?_⟩ + let f : ℕ → Ideal R := fun n ↦ (I.colon (span {b ^ n})) + have hf : Monotone f := by + intro n m hnm + simp_rw [f] + exact (Submodule.colon_mono le_rfl (Ideal.span_singleton_le_span_singleton.mpr + (pow_dvd_pow b hnm))) + obtain ⟨n, hn⟩ := monotone_stabilizes_iff_noetherian.mpr ‹_› ⟨f, hf⟩ + rcases h with ⟨-, h⟩ + specialize @h (I.colon (span {b ^ n})) (I + (span {b ^ n})) ?_ + · refine le_antisymm (fun r ↦ ?_) (le_inf (fun _ ↦ ?_) ?_) + · simp only [Submodule.add_eq_sup, sup_comm I, mem_inf, mem_colon_singleton, + mem_span_singleton_sup, and_imp, forall_exists_index] + rintro hrb t s hs rfl + refine add_mem ?_ hs + have := hn (n + n) (by simp) + simp only [OrderHom.coe_mk, f] at this + rw [add_mul, mul_assoc, ← pow_add] at hrb + rwa [← mem_colon_singleton, this, mem_colon_singleton, + ← Ideal.add_mem_iff_left _ (Ideal.mul_mem_right _ _ hs)] + · simpa only [mem_colon_singleton] using mul_mem_right _ _ + · simp + rcases h with (h|h) + · replace h : I = I.colon (span {b}) := by + rcases eq_or_ne n 0 with rfl|hn' + · simpa [f] using hn 1 zero_le_one + refine le_antisymm ?_ (h.le.trans' (Submodule.colon_mono le_rfl ?_)) + · intro + simpa only [mem_colon_singleton] using mul_mem_right _ _ + · exact span_singleton_le_span_singleton.mpr (dvd_pow_self b hn') + rw [← mem_colon_singleton, ← h] at hab + exact Or.inl hab + · rw [← h] + refine Or.inr ⟨n, ?_⟩ + simpa using mem_sup_right (mem_span_singleton_self _) + +lemma exists_isPrimary_decomposition (I : Ideal R) : + ∃ s : Finset (Ideal R), s.inf id = I ∧ ∀ ⦃J⦄, J ∈ s → J.IsPrimary := + (exists_infIrred_decomposition I).imp fun _ h ↦ h.imp_right fun h' _ ht ↦ (h' ht).isPrimary + +variable (R) + +/-- The Lasker--Noether theorem: every ideal in a Noetherian ring admits a decomposition into + primary ideals. -/ +lemma isLasker : IsLasker R := exists_isPrimary_decomposition + +end Noetherian + +end Ideal From 17cb78946c7baaf74ddd5d54debeb02ae3f6df5c Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Sun, 27 Oct 2024 15:34:23 +0000 Subject: [PATCH 476/505] feat(NumberTheory/Padics): p-adic valuation of casts (#17544) Lemmas about `Padic.valuation` applied to casts from various other rings to Qp --- Mathlib/NumberTheory/Padics/PadicNumbers.lean | 62 +++++++++++-------- 1 file changed, 35 insertions(+), 27 deletions(-) diff --git a/Mathlib/NumberTheory/Padics/PadicNumbers.lean b/Mathlib/NumberTheory/Padics/PadicNumbers.lean index 33559d1491dfe..d86fe1b155206 100644 --- a/Mathlib/NumberTheory/Padics/PadicNumbers.lean +++ b/Mathlib/NumberTheory/Padics/PadicNumbers.lean @@ -945,16 +945,6 @@ def valuation : ℚ_[p] → ℤ := theorem valuation_zero : valuation (0 : ℚ_[p]) = 0 := dif_pos ((const_equiv p).2 rfl) -@[simp] -theorem valuation_one : valuation (1 : ℚ_[p]) = 0 := by - classical - change dite (CauSeq.const (padicNorm p) 1 ≈ _) _ _ = _ - have h : ¬CauSeq.const (padicNorm p) 1 ≈ 0 := by - rw [← const_zero, const_equiv p] - exact one_ne_zero - rw [dif_neg h] - simp - theorem norm_eq_pow_val {x : ℚ_[p]} : x ≠ 0 → ‖x‖ = (p : ℝ) ^ (-x.valuation) := by refine Quotient.inductionOn' x fun f hf => ?_ change (PadicSeq.norm _ : ℝ) = (p : ℝ) ^ (-PadicSeq.valuation _) @@ -969,27 +959,45 @@ theorem norm_eq_pow_val {x : ℚ_[p]} : x ≠ 0 → ‖x‖ = (p : ℝ) ^ (-x.va simpa using hf' @[simp] -theorem valuation_p : valuation (p : ℚ_[p]) = 1 := by - have h : (1 : ℝ) < p := mod_cast (Fact.out : p.Prime).one_lt - refine neg_injective ((zpow_right_strictMono₀ h).injective <| (norm_eq_pow_val ?_).symm.trans ?_) - · exact mod_cast (Fact.out : p.Prime).ne_zero - · simp +lemma valuation_ratCast (q : ℚ) : valuation (q : ℚ_[p]) = padicValRat p q := by + rcases eq_or_ne q 0 with rfl | hq + · simp only [Rat.cast_zero, valuation_zero, padicValRat.zero] + refine neg_injective ((zpow_right_strictMono₀ (mod_cast hp.out.one_lt)).injective + <| (norm_eq_pow_val (mod_cast hq)).symm.trans ?_) + rw [padicNormE.eq_padicNorm, ← Rat.cast_natCast, ← Rat.cast_zpow, Rat.cast_inj] + exact padicNorm.eq_zpow_of_nonzero hq + +@[simp] +lemma valuation_intCast (n : ℤ) : valuation (n : ℚ_[p]) = padicValInt p n := by + rw [← Rat.cast_intCast, valuation_ratCast, padicValRat.of_int] + +@[simp] +lemma valuation_natCast (n : ℕ) : valuation (n : ℚ_[p]) = padicValNat p n := by + rw [← Rat.cast_natCast, valuation_ratCast, padicValRat.of_nat] + +-- See note [no_index around OfNat.ofNat] +@[simp] +lemma valuation_ofNat (n : ℕ) [n.AtLeastTwo] : + valuation (no_index (OfNat.ofNat n : ℚ_[p])) = padicValNat p n := + valuation_natCast n + +@[simp] +lemma valuation_one : valuation (1 : ℚ_[p]) = 0 := by + rw [← Nat.cast_one, valuation_natCast, padicValNat.one, cast_zero] + +-- not @[simp], since simp can prove it +lemma valuation_p : valuation (p : ℚ_[p]) = 1 := by + rw [valuation_natCast, padicValNat_self, cast_one] theorem valuation_map_add {x y : ℚ_[p]} (hxy : x + y ≠ 0) : min (valuation x) (valuation y) ≤ valuation (x + y : ℚ_[p]) := by by_cases hx : x = 0 - · rw [hx, zero_add] - exact min_le_right _ _ - · by_cases hy : y = 0 - · rw [hy, add_zero] - exact min_le_left _ _ - · have h_norm : ‖x + y‖ ≤ max ‖x‖ ‖y‖ := padicNormE.nonarchimedean x y - have hp_one : (1 : ℝ) < p := by - rw [← Nat.cast_one, Nat.cast_lt] - exact Nat.Prime.one_lt hp.elim - rwa [norm_eq_pow_val hx, norm_eq_pow_val hy, norm_eq_pow_val hxy, le_max_iff, - zpow_le_zpow_iff_right₀ hp_one, zpow_le_zpow_iff_right₀ hp_one, neg_le_neg_iff, - neg_le_neg_iff, ← min_le_iff] at h_norm + · simpa only [hx, zero_add] using min_le_right _ _ + by_cases hy : y = 0 + · simpa only [hy, add_zero] using min_le_left _ _ + have : ‖x + y‖ ≤ max ‖x‖ ‖y‖ := padicNormE.nonarchimedean x y + simpa only [norm_eq_pow_val hxy, norm_eq_pow_val hx, norm_eq_pow_val hy, le_max_iff, + zpow_le_zpow_iff_right₀ (mod_cast hp.out.one_lt : 1 < (p : ℝ)), neg_le_neg_iff, ← min_le_iff] @[simp] theorem valuation_map_mul {x y : ℚ_[p]} (hx : x ≠ 0) (hy : y ≠ 0) : From 5532a724f861e30fc9030573e712370d69a58579 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Sun, 27 Oct 2024 15:49:22 +0000 Subject: [PATCH 477/505] feat(CategoryTheory/SmallObject): very basic API for the iteration of functors (#18115) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR adds some very basic API for the iteration of functors and fixes the notation for intervals (`Set.Iic` and `Set.Iio` are now used). Co-authored-by: Joël Riou <37772949+joelriou@users.noreply.github.com> --- Mathlib.lean | 2 +- .../{Iteration.lean => Iteration/Basic.lean} | 104 +++++++++++++++--- 2 files changed, 90 insertions(+), 16 deletions(-) rename Mathlib/CategoryTheory/SmallObject/{Iteration.lean => Iteration/Basic.lean} (65%) diff --git a/Mathlib.lean b/Mathlib.lean index 46d19d8a3517b..93e4966d9d628 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1988,7 +1988,7 @@ import Mathlib.CategoryTheory.Sites.Types import Mathlib.CategoryTheory.Sites.Whiskering import Mathlib.CategoryTheory.Skeletal import Mathlib.CategoryTheory.SmallObject.Construction -import Mathlib.CategoryTheory.SmallObject.Iteration +import Mathlib.CategoryTheory.SmallObject.Iteration.Basic import Mathlib.CategoryTheory.Square import Mathlib.CategoryTheory.Subobject.Basic import Mathlib.CategoryTheory.Subobject.Comma diff --git a/Mathlib/CategoryTheory/SmallObject/Iteration.lean b/Mathlib/CategoryTheory/SmallObject/Iteration/Basic.lean similarity index 65% rename from Mathlib/CategoryTheory/SmallObject/Iteration.lean rename to Mathlib/CategoryTheory/SmallObject/Iteration/Basic.lean index f4d4af8823a9f..9bcc16aca6810 100644 --- a/Mathlib/CategoryTheory/SmallObject/Iteration.lean +++ b/Mathlib/CategoryTheory/SmallObject/Iteration/Basic.lean @@ -15,7 +15,7 @@ In this file, given a functor `Φ : C ⥤ C` and a natural transformation Given `j : J` where `J` is a well ordered set, we first introduce a category `Iteration ε j`. An object in this category consists of -a functor `F : { i // i ≤ j } ⥤ C ⥤ C` equipped with the data +a functor `F : Set.Iic j ⥤ C ⥤ C` equipped with the data which makes it the `i`th-iteration of `Φ` for all `i` such that `i ≤ j`. Under suitable assumptions on `C`, we shall show that this category `Iteration ε j` is equivalent to the punctual category (TODO). @@ -37,16 +37,15 @@ namespace CategoryTheory open Category Limits -variable {C : Type*} [Category C] {Φ : C ⥤ C} (ε : 𝟭 C ⟶ Φ) - {J : Type u} [Preorder J] +variable {C : Type*} [Category C] {Φ : C ⥤ C} (ε : 𝟭 C ⟶ Φ) {J : Type u} namespace Functor namespace Iteration -variable {j : J} (F : { i // i ≤ j } ⥤ C) +variable [Preorder J] {j : J} (F : Set.Iic j ⥤ C) -/-- The map `F.obj ⟨i, _⟩ ⟶ F.obj ⟨Order.succ i, _⟩` when `F : { i // i ≤ j } ⥤ C` +/-- The map `F.obj ⟨i, _⟩ ⟶ F.obj ⟨Order.succ i, _⟩` when `F : Set.Iic j ⥤ C` and `i : J` is such that `i < j`. -/ noncomputable abbrev mapSucc' [SuccOrder J] (i : J) (hi : i < j) : F.obj ⟨i, hi.le⟩ ⟶ F.obj ⟨Order.succ i, Order.succ_le_of_lt hi⟩ := @@ -54,9 +53,9 @@ noncomputable abbrev mapSucc' [SuccOrder J] (i : J) (hi : i < j) : variable {i : J} (hi : i ≤ j) -/-- The functor `{ k // k < i } ⥤ C` obtained by "restriction" of `F : { i // i ≤ j } ⥤ C` +/-- The functor `Set.Iio i ⥤ C` obtained by "restriction" of `F : Set.Iic j ⥤ C` when `i ≤ j`. -/ -def restrictionLT : { k // k < i } ⥤ C := +def restrictionLT : Set.Iio i ⥤ C := (monotone_inclusion_lt_le_of_le hi).functor ⋙ F @[simp] @@ -64,10 +63,10 @@ lemma restrictionLT_obj (k : J) (hk : k < i) : (restrictionLT F hi).obj ⟨k, hk⟩ = F.obj ⟨k, hk.le.trans hi⟩ := rfl @[simp] -lemma restrictionLT_map {k₁ k₂ : { k // k < i }} (φ : k₁ ⟶ k₂) : +lemma restrictionLT_map {k₁ k₂ : Set.Iio i} (φ : k₁ ⟶ k₂) : (restrictionLT F hi).map φ = F.map (homOfLE (by simpa using leOfHom φ)) := rfl -/-- Given `F : { i // i ≤ j } ⥤ C`, `i : J` such that `hi : i ≤ j`, this is the +/-- Given `F : Set.Iic j ⥤ C`, `i : J` such that `hi : i ≤ j`, this is the cocone consisting of all maps `F.obj ⟨k, hk⟩ ⟶ F.obj ⟨i, hi⟩` for `k : J` such that `k < i`. -/ @[simps] def coconeOfLE : Cocone (restrictionLT F hi) where @@ -77,6 +76,19 @@ def coconeOfLE : Cocone (restrictionLT F hi) where naturality := fun ⟨k₁, hk₁⟩ ⟨k₂, hk₂⟩ _ => by simp [comp_id, ← Functor.map_comp, homOfLE_comp] } +/-- The functor `Set.Iic i ⥤ C` obtained by "restriction" of `F : Set.Iic j ⥤ C` +when `i ≤ j`. -/ +def restrictionLE : Set.Iic i ⥤ C := + (monotone_inclusion_le_le_of_le hi).functor ⋙ F + +@[simp] +lemma restrictionLE_obj (k : J) (hk : k ≤ i) : + (restrictionLE F hi).obj ⟨k, hk⟩ = F.obj ⟨k, hk.trans hi⟩ := rfl + +@[simp] +lemma restrictionLE_map {k₁ k₂ : Set.Iic i} (φ : k₁ ⟶ k₂) : + (restrictionLE F hi).map φ = F.map (homOfLE (by simpa using leOfHom φ)) := rfl + end Iteration /-- The category of `j`th iterations of a functor `Φ` equipped with a natural @@ -86,7 +98,7 @@ equipped with data and properties which characterizes the iterations up to a uni isomorphism for the three types of elements: `⊥`, successors, limit elements. -/ structure Iteration [Preorder J] [OrderBot J] [SuccOrder J] (j : J) where /-- The data of all `i`th iterations for `i : J` such that `i ≤ j`. -/ - F : { i // i ≤ j } ⥤ C ⥤ C + F : Set.Iic j ⥤ C ⥤ C /-- The zeroth iteration is the identity functor. -/ isoZero : F.obj ⟨⊥, bot_le⟩ ≅ 𝟭 C /-- The iteration on a successor element is obtained by composition of @@ -107,7 +119,7 @@ variable {j : J} section -variable [OrderBot J] [SuccOrder J] (iter : Φ.Iteration ε j) +variable [Preorder J] [OrderBot J] [SuccOrder J] (iter : Φ.Iteration ε j) /-- For `iter : Φ.Iteration.ε j`, this is the map `iter.F.obj ⟨i, _⟩ ⟶ iter.F.obj ⟨Order.succ i, _⟩` if `i : J` is such that `i < j`. -/ @@ -121,7 +133,9 @@ lemma mapSucc_eq (i : J) (hi : i < j) : end -variable [OrderBot J] [SuccOrder J] (iter₁ iter₂ : Φ.Iteration ε j) +section + +variable [Preorder J] [OrderBot J] [SuccOrder J] (iter₁ iter₂ : Φ.Iteration ε j) /-- A morphism between two objects `iter₁` and `iter₂` in the category `Φ.Iteration ε j` of `j`th iterations of a functor `Φ` @@ -181,19 +195,79 @@ instance {J} {j : J} [PartialOrder J] [OrderBot J] [WellFoundedLT J] [SuccOrder intro i induction i using SuccOrder.limitRecOn with | hm j H => - have := H.eq_bot - subst this + obtain rfl := H.eq_bot simp [natTrans_app_zero] | hs j H IH => intro hj simp [Hom.natTrans_app_succ, IH, (Order.lt_succ_of_not_isMax H).trans_le hj] | hl j H IH => - apply fun hj ↦ (iter₁.isColimit j H hj).hom_ext ?_ + refine fun hj ↦ (iter₁.isColimit j H hj).hom_ext ?_ rintro ⟨k, hk⟩ simp [IH k hk] end Hom +@[simp] +lemma natTrans_id : Hom.natTrans (𝟙 iter₁) = 𝟙 _ := rfl + +variable {iter₁ iter₂} + +@[simp, reassoc] +lemma natTrans_comp {iter₃ : Iteration ε j} (φ : iter₁ ⟶ iter₂) (ψ : iter₂ ⟶ iter₃) : + (φ ≫ ψ).natTrans = φ.natTrans ≫ ψ.natTrans := rfl + +@[reassoc] +lemma natTrans_naturality (φ : iter₁ ⟶ iter₂) (i₁ i₂ : J) (h : i₁ ≤ i₂) (h' : i₂ ≤ j) : + iter₁.F.map (by exact homOfLE h) ≫ φ.natTrans.app ⟨i₂, h'⟩ = + φ.natTrans.app ⟨i₁, h.trans h'⟩ ≫ iter₂.F.map (by exact homOfLE h) := by + apply φ.natTrans.naturality + +variable (ε) in +/-- The evaluation functor `Iteration ε j ⥤ C ⥤ C` at `i : J` when `i ≤ j`. -/ +@[simps] +def eval {i : J} (hi : i ≤ j) : Iteration ε j ⥤ C ⥤ C where + obj iter := iter.F.obj ⟨i, hi⟩ + map φ := φ.natTrans.app _ + +/-- Given `iter : Iteration ε j` and `i : J` such that `i ≤ j`, this is the +induced element in `Iteration ε i`. -/ +@[simps F isoZero isoSucc] +def trunc (iter : Iteration ε j) {i : J} (hi : i ≤ j) : Iteration ε i where + F := restrictionLE (iter.F) hi + isoZero := iter.isoZero + isoSucc k hk := iter.isoSucc k (lt_of_lt_of_le hk hi) + mapSucc'_eq k hk := iter.mapSucc'_eq k (lt_of_lt_of_le hk hi) + isColimit k hk' hk := iter.isColimit k hk' (hk.trans hi) + +variable (ε) in +/-- The truncation functor `Iteration ε j ⥤ Iteration ε i` when `i ≤ j`. -/ +@[simps obj] +def truncFunctor {i : J} (hi : i ≤ j) : Iteration ε j ⥤ Iteration ε i where + obj iter := iter.trunc hi + map {iter₁ iter₂} φ := + { natTrans := whiskerLeft _ φ.natTrans + natTrans_app_succ := fun k hk => φ.natTrans_app_succ k (lt_of_lt_of_le hk hi) } + +@[simp] +lemma truncFunctor_map_natTrans_app + (φ : iter₁ ⟶ iter₂) {i : J} (hi : i ≤ j) (k : J) (hk : k ≤ i) : + ((truncFunctor ε hi).map φ).natTrans.app ⟨k, hk⟩ = + φ.natTrans.app ⟨k, hk.trans hi⟩ := rfl + +end + +namespace Hom + +variable [PartialOrder J] [OrderBot J] [SuccOrder J] [WellFoundedLT J] + (iter₁ iter₂ : Φ.Iteration ε j) + +lemma congr_app (φ φ' : iter₁ ⟶ iter₂) (i : J) (hi : i ≤ j) : + φ.natTrans.app ⟨i, hi⟩ = φ'.natTrans.app ⟨i, hi⟩ := by + obtain rfl := Subsingleton.elim φ φ' + rfl + +end Hom + end Iteration end Functor From 0311b12d9c874ba88246fea92c992b97025af76d Mon Sep 17 00:00:00 2001 From: "Tristan F.-R." Date: Sun, 27 Oct 2024 15:56:25 +0000 Subject: [PATCH 478/505] golf(Data/Complex/Abs): drop 1 != 0 (#18272) drops the instance `NeZero 1` as it is no longer needed for `Real.sqrt_mul two_pos.le` Co-authored-by: LeoDog896 --- Mathlib/Data/Complex/Abs.lean | 5 ----- Mathlib/LinearAlgebra/QuadraticForm/Real.lean | 1 - .../Kernel/Disintegration/MeasurableStieltjes.lean | 1 - 3 files changed, 7 deletions(-) diff --git a/Mathlib/Data/Complex/Abs.lean b/Mathlib/Data/Complex/Abs.lean index bd36ccce8e8a2..99319c9d298b3 100644 --- a/Mathlib/Data/Complex/Abs.lean +++ b/Mathlib/Data/Complex/Abs.lean @@ -170,11 +170,6 @@ theorem abs_abs (z : ℂ) : |Complex.abs z| = Complex.abs z := theorem abs_le_abs_re_add_abs_im (z : ℂ) : Complex.abs z ≤ |z.re| + |z.im| := by simpa [re_add_im] using Complex.abs.add_le z.re (z.im * I) --- Porting note: added so `two_pos` in the next proof works --- TODO: move somewhere else -instance : NeZero (1 : ℝ) := - ⟨by apply one_ne_zero⟩ - theorem abs_le_sqrt_two_mul_max (z : ℂ) : Complex.abs z ≤ Real.sqrt 2 * max |z.re| |z.im| := by cases' z with x y simp only [abs_apply, normSq_mk, ← sq] diff --git a/Mathlib/LinearAlgebra/QuadraticForm/Real.lean b/Mathlib/LinearAlgebra/QuadraticForm/Real.lean index b9fe9bdf90ae3..69811abb31c01 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/Real.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/Real.lean @@ -7,7 +7,6 @@ import Mathlib.LinearAlgebra.QuadraticForm.IsometryEquiv import Mathlib.Data.Sign import Mathlib.Algebra.CharP.Invertible import Mathlib.Analysis.RCLike.Basic -import Mathlib.Data.Complex.Abs /-! # Real quadratic forms diff --git a/Mathlib/Probability/Kernel/Disintegration/MeasurableStieltjes.lean b/Mathlib/Probability/Kernel/Disintegration/MeasurableStieltjes.lean index 3dfd2fa43ee37..251bef27aec44 100644 --- a/Mathlib/Probability/Kernel/Disintegration/MeasurableStieltjes.lean +++ b/Mathlib/Probability/Kernel/Disintegration/MeasurableStieltjes.lean @@ -3,7 +3,6 @@ Copyright (c) 2024 Rémy Degenne. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ -import Mathlib.Data.Complex.Abs import Mathlib.MeasureTheory.Measure.GiryMonad import Mathlib.MeasureTheory.Measure.Stieltjes import Mathlib.Analysis.Normed.Order.Lattice From 498e03b20ae8b8e5689e3eb047de1406eefeb47f Mon Sep 17 00:00:00 2001 From: Markus Himmel Date: Sun, 27 Oct 2024 16:44:05 +0000 Subject: [PATCH 479/505] chore: fix a Lean-3-ism in a module doc (#18289) --- Mathlib/CategoryTheory/Limits/Preserves/Opposites.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Opposites.lean b/Mathlib/CategoryTheory/Limits/Preserves/Opposites.lean index d56a29ad3be6d..323fb438c4e70 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Opposites.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Opposites.lean @@ -7,7 +7,7 @@ import Mathlib.CategoryTheory.Limits.Opposites import Mathlib.CategoryTheory.Limits.Preserves.Finite /-! -# Limit preservation properties of `functor.op` and related constructions +# Limit preservation properties of `Functor.op` and related constructions We formulate conditions about `F` which imply that `F.op`, `F.unop`, `F.leftOp` and `F.rightOp` preserve certain (co)limits. From 8c2c97e8beb66d75f3ddc8d4c6878cd4dca30f2d Mon Sep 17 00:00:00 2001 From: Markus Himmel Date: Sun, 27 Oct 2024 16:51:19 +0000 Subject: [PATCH 480/505] chore: fix lemma names for `preservesLimitIso` (#18288) --- .../Category/ModuleCat/Presheaf/Colimits.lean | 10 +++--- .../Category/ModuleCat/Presheaf/Limits.lean | 10 +++--- Mathlib/CategoryTheory/GlueData.lean | 2 +- Mathlib/CategoryTheory/Limits/Elements.lean | 4 +-- Mathlib/CategoryTheory/Limits/IndYoneda.lean | 4 +-- .../Limits/Preserves/Limits.lean | 36 +++++++++++++------ .../Limits/Preserves/Yoneda.lean | 2 +- Mathlib/Geometry/Manifold/Sheaf/Smooth.lean | 2 +- .../Geometry/RingedSpace/OpenImmersion.lean | 12 +++---- .../PresheafedSpace/HasColimits.lean | 2 +- Mathlib/Topology/Gluing.lean | 2 +- 11 files changed, 51 insertions(+), 35 deletions(-) diff --git a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Colimits.lean b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Colimits.lean index ea25e5d1bf2f7..f58ed17c074ed 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Colimits.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Colimits.lean @@ -68,17 +68,17 @@ noncomputable def colimitPresheafOfModules : PresheafOfModules R where map_id X := colimit.hom_ext (fun j => by dsimp rw [ι_colimMap_assoc, whiskerLeft_app, restriction_app] - erw [ι_preservesColimitsIso_inv (G := ModuleCat.restrictScalars (R.map (𝟙 X))), + erw [ι_preservesColimitIso_inv (G := ModuleCat.restrictScalars (R.map (𝟙 X))), ModuleCat.restrictScalarsId'App_inv_naturality] rw [map_id] dsimp) map_comp {X Y Z} f g := colimit.hom_ext (fun j => by dsimp rw [ι_colimMap_assoc, whiskerLeft_app, restriction_app, assoc, ι_colimMap_assoc] - erw [ι_preservesColimitsIso_inv (G := ModuleCat.restrictScalars (R.map (f ≫ g))), - ι_preservesColimitsIso_inv_assoc (G := ModuleCat.restrictScalars (R.map f))] + erw [ι_preservesColimitIso_inv (G := ModuleCat.restrictScalars (R.map (f ≫ g))), + ι_preservesColimitIso_inv_assoc (G := ModuleCat.restrictScalars (R.map f))] rw [← Functor.map_comp_assoc, ι_colimMap_assoc] - erw [ι_preservesColimitsIso_inv (G := ModuleCat.restrictScalars (R.map g))] + erw [ι_preservesColimitIso_inv (G := ModuleCat.restrictScalars (R.map g))] rw [map_comp, ModuleCat.restrictScalarsComp'_inv_app, assoc, assoc, whiskerLeft_app, whiskerLeft_app, restriction_app, restriction_app] simp only [Functor.map_comp, assoc] @@ -94,7 +94,7 @@ noncomputable def colimitCocone : Cocone F where { app := fun X ↦ colimit.ι (F ⋙ evaluation R X) j naturality := fun {X Y} f ↦ by dsimp - erw [colimit.ι_desc_assoc, assoc, ← ι_preservesColimitsIso_inv] + erw [colimit.ι_desc_assoc, assoc, ← ι_preservesColimitIso_inv] rfl } naturality := fun {X Y} f ↦ by ext1 X diff --git a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Limits.lean b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Limits.lean index 52757e35c210f..127f29b50a89e 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Limits.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Limits.lean @@ -73,7 +73,7 @@ noncomputable def limitPresheafOfModules : PresheafOfModules R where dsimp simp only [limMap_π, Functor.comp_obj, evaluation_obj, whiskerLeft_app, restriction_app, assoc] - erw [preservesLimitsIso_hom_π] + erw [preservesLimitIso_hom_π] rw [← ModuleCat.restrictScalarsId'App_inv_naturality, map_id, ModuleCat.restrictScalarsId'_inv_app] dsimp @@ -85,14 +85,14 @@ noncomputable def limitPresheafOfModules : PresheafOfModules R where intro j simp only [Functor.comp_obj, evaluation_obj, limMap_π, whiskerLeft_app, restriction_app, map_comp, ModuleCat.restrictScalarsComp'_inv_app, Functor.map_comp, assoc] - erw [preservesLimitsIso_hom_π] + erw [preservesLimitIso_hom_π] rw [← ModuleCat.restrictScalarsComp'App_inv_naturality] dsimp rw [← Functor.map_comp_assoc, ← Functor.map_comp_assoc, assoc, - preservesLimitsIso_inv_π] + preservesLimitIso_inv_π] erw [limMap_π] dsimp - simp only [Functor.map_comp, assoc, preservesLimitsIso_inv_π_assoc] + simp only [Functor.map_comp, assoc, preservesLimitIso_inv_π_assoc] erw [limMap_π_assoc] dsimp @@ -106,7 +106,7 @@ noncomputable def limitCone : Cone F where { app := fun X ↦ limit.π (F ⋙ evaluation R X) j naturality := fun {X Y} f ↦ by dsimp - simp only [assoc, preservesLimitsIso_inv_π] + simp only [assoc, preservesLimitIso_inv_π] apply limMap_π } naturality := fun {j j'} f ↦ by ext1 X diff --git a/Mathlib/CategoryTheory/GlueData.lean b/Mathlib/CategoryTheory/GlueData.lean index b43f3893e3bca..433fa5d3d835c 100644 --- a/Mathlib/CategoryTheory/GlueData.lean +++ b/Mathlib/CategoryTheory/GlueData.lean @@ -318,7 +318,7 @@ def gluedIso : F.obj D.glued ≅ (D.mapGlueData F).glued := @[reassoc (attr := simp)] theorem ι_gluedIso_hom (i : D.J) : F.map (D.ι i) ≫ (D.gluedIso F).hom = (D.mapGlueData F).ι i := by haveI : HasColimit (MultispanIndex.multispan (diagram (mapGlueData D F))) := inferInstance - erw [ι_preservesColimitsIso_hom_assoc] + erw [ι_preservesColimitIso_hom_assoc] rw [HasColimit.isoOfNatIso_ι_hom] erw [Category.id_comp] rfl diff --git a/Mathlib/CategoryTheory/Limits/Elements.lean b/Mathlib/CategoryTheory/Limits/Elements.lean index 84963f2bacdff..bea7b175e4f22 100644 --- a/Mathlib/CategoryTheory/Limits/Elements.lean +++ b/Mathlib/CategoryTheory/Limits/Elements.lean @@ -53,7 +53,7 @@ lemma map_lift_mapCone (c : Cone F) : A.map (limit.lift (F ⋙ π A) ((π A).mapCone c)) c.pt.snd = liftedConeElement F := by apply (preservesLimitIso A (F ⋙ π A)).toEquiv.injective ext i - have h₁ := congrFun (preservesLimitsIso_hom_π A (F ⋙ π A) i) + have h₁ := congrFun (preservesLimitIso_hom_π A (F ⋙ π A) i) (A.map (limit.lift (F ⋙ π A) ((π A).mapCone c)) c.pt.snd) have h₂ := (c.π.app i).property simp_all [← FunctorToTypes.map_comp_apply, liftedConeElement] @@ -62,7 +62,7 @@ lemma map_lift_mapCone (c : Cone F) : lemma map_π_liftedConeElement (i : I) : A.map (limit.π (F ⋙ π A) i) (liftedConeElement F) = (F.obj i).snd := by have := congrFun - (preservesLimitsIso_inv_π A (F ⋙ π A) i) (liftedConeElement' F) + (preservesLimitIso_inv_π A (F ⋙ π A) i) (liftedConeElement' F) simp_all [liftedConeElement] /-- (implementation) The constructured limit cone. -/ diff --git a/Mathlib/CategoryTheory/Limits/IndYoneda.lean b/Mathlib/CategoryTheory/Limits/IndYoneda.lean index 99842d2a0d02f..07f8ecd466bc2 100644 --- a/Mathlib/CategoryTheory/Limits/IndYoneda.lean +++ b/Mathlib/CategoryTheory/Limits/IndYoneda.lean @@ -94,7 +94,7 @@ lemma coyonedaOpColimitIsoLimitCoyoneda_hom_comp_π (i : I) : (coyonedaOpColimitIsoLimitCoyoneda F).hom ≫ limit.π (F.op.comp coyoneda) ⟨i⟩ = coyoneda.map (colimit.ι F i).op := by simp only [coyonedaOpColimitIsoLimitCoyoneda, Functor.mapIso_symm, - Iso.trans_hom, Iso.symm_hom, Functor.mapIso_inv, Category.assoc, preservesLimitsIso_hom_π, + Iso.trans_hom, Iso.symm_hom, Functor.mapIso_inv, Category.assoc, preservesLimitIso_hom_π, ← Functor.map_comp, limitOpIsoOpColimit_inv_comp_π] @[reassoc (attr := simp)] @@ -142,7 +142,7 @@ lemma coyonedaOpColimitIsoLimitCoyoneda'_hom_comp_π (i : I) : (coyonedaOpColimitIsoLimitCoyoneda' F).hom ≫ limit.π (F.rightOp ⋙ coyoneda) i = coyoneda.map (colimit.ι F ⟨i⟩).op := by simp only [coyonedaOpColimitIsoLimitCoyoneda', Functor.mapIso_symm, Iso.trans_hom, Iso.symm_hom, - Functor.mapIso_inv, Category.assoc, preservesLimitsIso_hom_π, ← Functor.map_comp, + Functor.mapIso_inv, Category.assoc, preservesLimitIso_hom_π, ← Functor.map_comp, limitRightOpIsoOpColimit_inv_comp_π] @[reassoc (attr := simp)] diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Limits.lean b/Mathlib/CategoryTheory/Limits/Preserves/Limits.lean index 10c785ddbe059..ac77b2c0bcd51 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Limits.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Limits.lean @@ -50,22 +50,29 @@ def preservesLimitIso : G.obj (limit F) ≅ limit (F ⋙ G) := (PreservesLimit.preserves (limit.isLimit _)).conePointUniqueUpToIso (limit.isLimit _) @[reassoc (attr := simp)] -theorem preservesLimitsIso_hom_π (j) : +theorem preservesLimitIso_hom_π (j) : (preservesLimitIso G F).hom ≫ limit.π _ j = G.map (limit.π F j) := IsLimit.conePointUniqueUpToIso_hom_comp _ _ j +@[deprecated (since := "2024-10-27")] alias preservesLimitsIso_hom_π := preservesLimitIso_hom_π + @[reassoc (attr := simp)] -theorem preservesLimitsIso_inv_π (j) : +theorem preservesLimitIso_inv_π (j) : (preservesLimitIso G F).inv ≫ G.map (limit.π F j) = limit.π _ j := IsLimit.conePointUniqueUpToIso_inv_comp _ _ j +@[deprecated (since := "2024-10-27")] alias preservesLimitsIso_inv_π := preservesLimitIso_inv_π + @[reassoc (attr := simp)] -theorem lift_comp_preservesLimitsIso_hom (t : Cone F) : +theorem lift_comp_preservesLimitIso_hom (t : Cone F) : G.map (limit.lift _ t) ≫ (preservesLimitIso G F).hom = limit.lift (F ⋙ G) (G.mapCone _) := by ext simp [← G.map_comp] +@[deprecated (since := "2024-10-27")] +alias lift_comp_preservesLimitsIso_hom := lift_comp_preservesLimitIso_hom + instance : IsIso (limit.post F G) := show IsIso (preservesLimitIso G F).hom from inferInstance @@ -80,8 +87,8 @@ def preservesLimitNatIso : lim ⋙ G ≅ (whiskeringRight J C D).obj G ⋙ lim : intro _ _ f apply limit.hom_ext; intro j dsimp - simp only [preservesLimitsIso_hom_π, whiskerRight_app, limMap_π, Category.assoc, - preservesLimitsIso_hom_π_assoc, ← G.map_comp]) + simp only [preservesLimitIso_hom_π, whiskerRight_app, limMap_π, Category.assoc, + preservesLimitIso_hom_π_assoc, ← G.map_comp]) end @@ -117,22 +124,31 @@ def preservesColimitIso : G.obj (colimit F) ≅ colimit (F ⋙ G) := (PreservesColimit.preserves (colimit.isColimit _)).coconePointUniqueUpToIso (colimit.isColimit _) @[reassoc (attr := simp)] -theorem ι_preservesColimitsIso_inv (j : J) : +theorem ι_preservesColimitIso_inv (j : J) : colimit.ι _ j ≫ (preservesColimitIso G F).inv = G.map (colimit.ι F j) := IsColimit.comp_coconePointUniqueUpToIso_inv _ (colimit.isColimit (F ⋙ G)) j +@[deprecated (since := "2024-10-27")] +alias ι_preservesColimitsIso_inv := ι_preservesColimitIso_inv + @[reassoc (attr := simp)] -theorem ι_preservesColimitsIso_hom (j : J) : +theorem ι_preservesColimitIso_hom (j : J) : G.map (colimit.ι F j) ≫ (preservesColimitIso G F).hom = colimit.ι (F ⋙ G) j := (PreservesColimit.preserves (colimit.isColimit _)).comp_coconePointUniqueUpToIso_hom _ j +@[deprecated (since := "2024-10-27")] +alias ι_preservesColimitsIso_hom := ι_preservesColimitIso_hom + @[reassoc (attr := simp)] -theorem preservesColimitsIso_inv_comp_desc (t : Cocone F) : +theorem preservesColimitIso_inv_comp_desc (t : Cocone F) : (preservesColimitIso G F).inv ≫ G.map (colimit.desc _ t) = colimit.desc _ (G.mapCocone t) := by ext simp [← G.map_comp] +@[deprecated (since := "2024-10-27")] +alias preservesColimitsIso_inv_comp_desc := preservesColimitIso_inv_comp_desc + instance : IsIso (colimit.post F G) := show IsIso (preservesColimitIso G F).inv from inferInstance @@ -149,8 +165,8 @@ def preservesColimitNatIso : colim ⋙ G ≅ (whiskeringRight J C D).obj G ⋙ c apply colimit.hom_ext; intro j dsimp rw [ι_colimMap_assoc] - simp only [ι_preservesColimitsIso_inv, whiskerRight_app, Category.assoc, - ι_preservesColimitsIso_inv_assoc, ← G.map_comp] + simp only [ι_preservesColimitIso_inv, whiskerRight_app, Category.assoc, + ι_preservesColimitIso_inv_assoc, ← G.map_comp] rw [ι_colimMap]) end diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Yoneda.lean b/Mathlib/CategoryTheory/Limits/Preserves/Yoneda.lean index e0da3d65a8d90..dabe16ad73949 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Yoneda.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Yoneda.lean @@ -61,7 +61,7 @@ theorem yonedaYonedaColimit_app_inv {X : C} : ((yonedaYonedaColimit F).app (op X intro j rw [colimit.ι_post, ι_colimMap_assoc] simp only [← CategoryTheory.Functor.assoc, comp_evaluation] - rw [ι_preservesColimitsIso_inv_assoc, ← Functor.map_comp_assoc] + rw [ι_preservesColimitIso_inv_assoc, ← Functor.map_comp_assoc] simp only [← comp_evaluation] rw [colimitObjIsoColimitCompEvaluation_ι_inv, whiskerLeft_app] ext η Y f diff --git a/Mathlib/Geometry/Manifold/Sheaf/Smooth.lean b/Mathlib/Geometry/Manifold/Sheaf/Smooth.lean index 6cc1c879d0531..1bb560ec583ab 100644 --- a/Mathlib/Geometry/Manifold/Sheaf/Smooth.lean +++ b/Mathlib/Geometry/Manifold/Sheaf/Smooth.lean @@ -298,7 +298,7 @@ def smoothSheafCommRing.forgetStalk (x : TopCat.of M) : (colimit.ι ((OpenNhds.inclusion x).op ⋙ (smoothSheafCommRing IM I M R).presheaf) U)) (forgetStalk IM I M R x).hom = colimit.ι ((OpenNhds.inclusion x).op ⋙ (smoothSheaf IM I M R).presheaf) U := - ι_preservesColimitsIso_hom _ _ _ + ι_preservesColimitIso_hom _ _ _ @[simp, reassoc, elementwise] lemma smoothSheafCommRing.ι_forgetStalk_inv (x : TopCat.of M) (U) : colimit.ι ((OpenNhds.inclusion x).op ⋙ (smoothSheaf IM I M R).presheaf) U ≫ diff --git a/Mathlib/Geometry/RingedSpace/OpenImmersion.lean b/Mathlib/Geometry/RingedSpace/OpenImmersion.lean index f549a8e9ee841..6befd6aab6a80 100644 --- a/Mathlib/Geometry/RingedSpace/OpenImmersion.lean +++ b/Mathlib/Geometry/RingedSpace/OpenImmersion.lean @@ -688,7 +688,7 @@ instance sheafedSpace_hasPullback_of_right : HasPullback g f := instance sheafedSpace_pullback_snd_of_left : SheafedSpace.IsOpenImmersion (pullback.snd f g) := by delta pullback.snd - have : _ = limit.π (cospan f g) right := preservesLimitsIso_hom_π forget (cospan f g) right + have : _ = limit.π (cospan f g) right := preservesLimitIso_hom_π forget (cospan f g) right rw [← this] have := HasLimit.isoOfNatIso_hom_π (diagramIsoCospan (cospan f g ⋙ forget)) right erw [Category.comp_id] at this @@ -699,7 +699,7 @@ instance sheafedSpace_pullback_snd_of_left : instance sheafedSpace_pullback_fst_of_right : SheafedSpace.IsOpenImmersion (pullback.fst g f) := by delta pullback.fst - have : _ = limit.π (cospan g f) left := preservesLimitsIso_hom_π forget (cospan g f) left + have : _ = limit.π (cospan g f) left := preservesLimitIso_hom_π forget (cospan g f) left rw [← this] have := HasLimit.isoOfNatIso_hom_π (diagramIsoCospan (cospan g f ⋙ forget)) left erw [Category.comp_id] at this @@ -840,7 +840,7 @@ variable [HasLimits C] {ι : Type v} (F : Discrete ι ⥤ SheafedSpace.{_, v, v} (i : Discrete ι) theorem sigma_ι_isOpenEmbedding : IsOpenEmbedding (colimit.ι F i).base := by - rw [← show _ = (colimit.ι F i).base from ι_preservesColimitsIso_inv (SheafedSpace.forget C) F i] + rw [← show _ = (colimit.ι F i).base from ι_preservesColimitIso_inv (SheafedSpace.forget C) F i] have : _ = _ ≫ colimit.ι (Discrete.functor ((F ⋙ SheafedSpace.forget C).obj ∘ Discrete.mk)) i := HasColimit.isoOfNatIso_ι_hom Discrete.natIsoFunctor i rw [← Iso.eq_comp_inv] at this @@ -871,12 +871,12 @@ theorem image_preimage_is_empty (j : Discrete ι) (h : i ≠ j) (U : Opens (F.ob replace eq := ConcreteCategory.congr_arg (preservesColimitIso (SheafedSpace.forget C) F ≪≫ HasColimit.isoOfNatIso Discrete.natIsoFunctor ≪≫ TopCat.sigmaIsoSigma.{v, v} _).hom eq simp_rw [CategoryTheory.Iso.trans_hom, ← TopCat.comp_app, ← PresheafedSpace.comp_base] at eq - rw [ι_preservesColimitsIso_inv] at eq + rw [ι_preservesColimitIso_inv] at eq change ((SheafedSpace.forget C).map (colimit.ι F i) ≫ _) y = ((SheafedSpace.forget C).map (colimit.ι F j) ≫ _) x at eq cases i; cases j - rw [ι_preservesColimitsIso_hom_assoc, ι_preservesColimitsIso_hom_assoc, + rw [ι_preservesColimitIso_hom_assoc, ι_preservesColimitIso_hom_assoc, HasColimit.isoOfNatIso_ι_hom_assoc, HasColimit.isoOfNatIso_ι_hom_assoc, TopCat.sigmaIsoSigma_hom_ι, TopCat.sigmaIsoSigma_hom_ι] at eq exact h (congr_arg Discrete.mk (congr_arg Sigma.fst eq)) @@ -886,7 +886,7 @@ instance sigma_ι_isOpenImmersion [HasStrictTerminalObjects C] : base_open := sigma_ι_isOpenEmbedding F i c_iso U := by have e : colimit.ι F i = _ := - (ι_preservesColimitsIso_inv SheafedSpace.forgetToPresheafedSpace F i).symm + (ι_preservesColimitIso_inv SheafedSpace.forgetToPresheafedSpace F i).symm have H : IsOpenEmbedding (colimit.ι (F ⋙ SheafedSpace.forgetToPresheafedSpace) i ≫ diff --git a/Mathlib/Geometry/RingedSpace/PresheafedSpace/HasColimits.lean b/Mathlib/Geometry/RingedSpace/PresheafedSpace/HasColimits.lean index 9b2339708fe4b..6f94368b8f4a6 100644 --- a/Mathlib/Geometry/RingedSpace/PresheafedSpace/HasColimits.lean +++ b/Mathlib/Geometry/RingedSpace/PresheafedSpace/HasColimits.lean @@ -307,7 +307,7 @@ def colimitPresheafObjIsoComponentwiseLimit (F : J ⥤ PresheafedSpace.{_, _, v} refine congr_arg (Set.preimage · U.1) (funext fun x => ?_) erw [← TopCat.comp_app] congr - exact ι_preservesColimitsIso_inv (forget C) F (unop X) + exact ι_preservesColimitIso_inv (forget C) F (unop X) · intro X Y f change ((F.map f.unop).c.app _ ≫ _ ≫ _) ≫ (F.obj (unop Y)).presheaf.map _ = _ ≫ _ rw [TopCat.Presheaf.Pushforward.comp_inv_app] diff --git a/Mathlib/Topology/Gluing.lean b/Mathlib/Topology/Gluing.lean index 87c5a52e0a494..644c93f2a65fb 100644 --- a/Mathlib/Topology/Gluing.lean +++ b/Mathlib/Topology/Gluing.lean @@ -166,7 +166,7 @@ theorem eqvGen_of_π_eq let diagram := parallelPair 𝖣.diagram.fstSigmaMap 𝖣.diagram.sndSigmaMap ⋙ forget _ have : colimit.ι diagram one x = colimit.ι diagram one y := by dsimp only [coequalizer.π, ContinuousMap.toFun_eq_coe] at h - rw [← ι_preservesColimitsIso_hom, forget_map_eq_coe, types_comp_apply, h] + rw [← ι_preservesColimitIso_hom, forget_map_eq_coe, types_comp_apply, h] simp have : (colimit.ι diagram _ ≫ colim.map _ ≫ (colimit.isoColimitCocone _).hom) _ = From d91cd4279b4fec91a17a07c01488515ddfca759f Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Sun, 27 Oct 2024 18:23:39 +0000 Subject: [PATCH 481/505] =?UTF-8?q?feat:=20define=20the=20positive=20and?= =?UTF-8?q?=20negative=20parts=20of=20elements=20in=20a=20C=E2=8B=86-algeb?= =?UTF-8?q?ra=20(#18103)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- Mathlib.lean | 1 + Mathlib/Algebra/Algebra/Quasispectrum.lean | 5 +- .../ContinuousFunctionalCalculus/PosPart.lean | 132 ++++++++++++++++++ Mathlib/Analysis/Normed/Order/Lattice.lean | 2 + 4 files changed, 138 insertions(+), 2 deletions(-) create mode 100644 Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/PosPart.lean diff --git a/Mathlib.lean b/Mathlib.lean index 93e4966d9d628..f176ac0021cf7 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1007,6 +1007,7 @@ import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Instances import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Integral import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.NonUnital import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Order +import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.PosPart import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Restrict import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Unique import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Unital diff --git a/Mathlib/Algebra/Algebra/Quasispectrum.lean b/Mathlib/Algebra/Algebra/Quasispectrum.lean index e94372dc43fdc..63ac30ba8df9a 100644 --- a/Mathlib/Algebra/Algebra/Quasispectrum.lean +++ b/Mathlib/Algebra/Algebra/Quasispectrum.lean @@ -341,8 +341,9 @@ lemma mem_spectrum_inr_of_not_isUnit {R A : Type*} [CommRing R] (a : A) (r : R) (hr : ¬ IsUnit r) : r ∈ spectrum R (a : Unitization R A) := fun h ↦ hr <| by simpa [map_sub] using h.map (fstHom R A) -lemma quasispectrum_eq_spectrum_inr (R : Type*) {A : Type*} [CommRing R] [Ring A] - [Algebra R A] (a : A) : quasispectrum R a = spectrum R (a : Unitization R A) := by +lemma quasispectrum_eq_spectrum_inr (R : Type*) {A : Type*} [CommRing R] [NonUnitalRing A] + [Module R A] [IsScalarTower R A A] [SMulCommClass R A A] (a : A) : + quasispectrum R a = spectrum R (a : Unitization R A) := by ext r have : { r | ¬ IsUnit r} ⊆ spectrum R _ := mem_spectrum_inr_of_not_isUnit a rw [← Set.union_eq_left.mpr this, ← quasispectrum_eq_spectrum_union] diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/PosPart.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/PosPart.lean new file mode 100644 index 0000000000000..dbdc386d6354d --- /dev/null +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/PosPart.lean @@ -0,0 +1,132 @@ +/- +Copyright (c) 2024 Jireh Loreaux. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jireh Loreaux +-/ + +import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.NonUnital +import Mathlib.Topology.ContinuousMap.StarOrdered + +/-! # The positive (and negative) parts of a selfadjoint element in a C⋆-algebra + +This file defines the positive and negative parts of a selfadjoint element in a C⋆-algebra via +the continuous functional calculus and develops the basic API. + +## TODO + ++ Prove that the positive and negative parts are unique. + +-/ + +variable {A : Type*} [NonUnitalRing A] [Module ℝ A] [SMulCommClass ℝ A A] [IsScalarTower ℝ A A] +variable [StarRing A] [TopologicalSpace A] +variable [NonUnitalContinuousFunctionalCalculus ℝ (IsSelfAdjoint : A → Prop)] + +namespace CStarAlgebra + +noncomputable instance : PosPart A where + posPart := cfcₙ (·⁺ : ℝ → ℝ) + +noncomputable instance : NegPart A where + negPart := cfcₙ (·⁻ : ℝ → ℝ) + +end CStarAlgebra + +namespace CFC + +lemma posPart_def (a : A) : a⁺ = cfcₙ (·⁺ : ℝ → ℝ) a := rfl + +lemma negPart_def (a : A) : a⁻ = cfcₙ (·⁻ : ℝ → ℝ) a := rfl + +@[simp] +lemma posPart_mul_negPart (a : A) : a⁺ * a⁻ = 0 := by + rw [posPart_def, negPart_def] + by_cases ha : IsSelfAdjoint a + · rw [← cfcₙ_mul _ _, ← cfcₙ_zero ℝ a] + refine cfcₙ_congr (fun x _ ↦ ?_) + simp only [_root_.posPart_def, _root_.negPart_def] + simpa using le_total x 0 + · simp [cfcₙ_apply_of_not_predicate a ha] + +@[simp] +lemma negPart_mul_posPart (a : A) : a⁻ * a⁺ = 0 := by + rw [posPart_def, negPart_def] + by_cases ha : IsSelfAdjoint a + · rw [← cfcₙ_mul _ _, ← cfcₙ_zero ℝ a] + refine cfcₙ_congr (fun x _ ↦ ?_) + simp only [_root_.posPart_def, _root_.negPart_def] + simpa using le_total 0 x + · simp [cfcₙ_apply_of_not_predicate a ha] + +lemma posPart_sub_negPart (a : A) (ha : IsSelfAdjoint a := by cfc_tac) : a⁺ - a⁻ = a := by + rw [posPart_def, negPart_def] + rw [← cfcₙ_sub _ _] + conv_rhs => rw [← cfcₙ_id ℝ a] + congr! 2 with + exact _root_.posPart_sub_negPart _ + +section Unique + +variable [UniqueNonUnitalContinuousFunctionalCalculus ℝ A] + +@[simp] +lemma posPart_neg (a : A) : (-a)⁺ = a⁻ := by + by_cases ha : IsSelfAdjoint a + · rw [posPart_def, negPart_def, ← cfcₙ_comp_neg _ _] + congr! 2 + · have ha' : ¬ IsSelfAdjoint (-a) := fun h ↦ ha (by simpa using h.neg) + rw [posPart_def, negPart_def, cfcₙ_apply_of_not_predicate a ha, + cfcₙ_apply_of_not_predicate _ ha'] + +@[simp] +lemma negPart_neg (a : A) : (-a)⁻ = a⁺ := by + rw [← eq_comm, ← sub_eq_zero, ← posPart_neg, neg_neg, sub_self] + +end Unique + +variable [PartialOrder A] [StarOrderedRing A] + +@[aesop norm apply (rule_sets := [CStarAlgebra])] +lemma posPart_nonneg (a : A) : + 0 ≤ a⁺ := + cfcₙ_nonneg (fun x _ ↦ by positivity) + +@[aesop norm apply (rule_sets := [CStarAlgebra])] +lemma negPart_nonneg (a : A) : + 0 ≤ a⁻ := + cfcₙ_nonneg (fun x _ ↦ by positivity) + +variable [NonnegSpectrumClass ℝ A] + +lemma eq_posPart_iff (a : A) : a = a⁺ ↔ 0 ≤ a := by + refine ⟨fun ha ↦ ha ▸ posPart_nonneg a, fun ha ↦ ?_⟩ + conv_lhs => rw [← cfcₙ_id ℝ a] + rw [posPart_def] + refine cfcₙ_congr (fun x hx ↦ ?_) + simpa [_root_.posPart_def] using quasispectrum_nonneg_of_nonneg a ha x hx + +lemma negPart_eq_zero_iff (a : A) (ha : IsSelfAdjoint a) : + a⁻ = 0 ↔ 0 ≤ a := by + rw [← eq_posPart_iff] + nth_rw 2 [← posPart_sub_negPart a] + simp + +lemma eq_negPart_iff (a : A) : a = -a⁻ ↔ a ≤ 0 := by + refine ⟨fun ha ↦ by rw [ha, neg_nonpos]; exact negPart_nonneg a, fun ha ↦ ?_⟩ + rw [← neg_nonneg] at ha + rw [negPart_def, ← cfcₙ_neg] + have _ : IsSelfAdjoint a := neg_neg a ▸ (IsSelfAdjoint.neg <| .of_nonneg ha) + conv_lhs => rw [← cfcₙ_id ℝ a] + refine cfcₙ_congr fun x hx ↦ ?_ + rw [Unitization.quasispectrum_eq_spectrum_inr ℝ, ← neg_neg x, ← Set.mem_neg, + spectrum.neg_eq, ← Unitization.inr_neg, ← Unitization.quasispectrum_eq_spectrum_inr ℝ] at hx + rw [← neg_eq_iff_eq_neg, eq_comm] + simpa using quasispectrum_nonneg_of_nonneg _ ha _ hx + +lemma posPart_eq_zero_iff (a : A) (ha : IsSelfAdjoint a) : + a⁺ = 0 ↔ a ≤ 0 := by + rw [← eq_negPart_iff] + nth_rw 2 [← posPart_sub_negPart a] + simp + +end CFC diff --git a/Mathlib/Analysis/Normed/Order/Lattice.lean b/Mathlib/Analysis/Normed/Order/Lattice.lean index b9217fe0951e2..de21daec89f71 100644 --- a/Mathlib/Analysis/Normed/Order/Lattice.lean +++ b/Mathlib/Analysis/Normed/Order/Lattice.lean @@ -185,8 +185,10 @@ lemma lipschitzWith_posPart : LipschitzWith 1 (posPart : α → α) := lemma lipschitzWith_negPart : LipschitzWith 1 (negPart : α → α) := by simpa [Function.comp] using lipschitzWith_posPart.comp LipschitzWith.id.neg +@[fun_prop] lemma continuous_posPart : Continuous (posPart : α → α) := lipschitzWith_posPart.continuous +@[fun_prop] lemma continuous_negPart : Continuous (negPart : α → α) := lipschitzWith_negPart.continuous lemma isClosed_nonneg : IsClosed {x : α | 0 ≤ x} := by From f0c67a9c1b8352739a65e8438f218cb11278ab12 Mon Sep 17 00:00:00 2001 From: Markus Himmel Date: Sun, 27 Oct 2024 18:47:18 +0000 Subject: [PATCH 482/505] feat: relate whiskering to evaluation (#18280) --- Mathlib/CategoryTheory/Products/Basic.lean | 26 ++++++++++++++++++++++ 1 file changed, 26 insertions(+) diff --git a/Mathlib/CategoryTheory/Products/Basic.lean b/Mathlib/CategoryTheory/Products/Basic.lean index dac0c3fe043ce..3abbf24a0925d 100644 --- a/Mathlib/CategoryTheory/Products/Basic.lean +++ b/Mathlib/CategoryTheory/Products/Basic.lean @@ -313,6 +313,32 @@ def compEvaluation (F : A ⥤ B ⥤ C) (b) : F ⋙ (evaluation _ _).obj b ≅ F. theorem comp_evaluation (F : A ⥤ B ⥤ C) (b) : F ⋙ (evaluation _ _).obj b = F.flip.obj b := rfl +/-- Whiskering by `F` and then evaluating at `a` is the same as evaluating at `F.obj a`. -/ +@[simps!] +def whiskeringLeftCompEvaluation (F : A ⥤ B) (a : A) : + (whiskeringLeft A B C).obj F ⋙ (evaluation A C).obj a ≅ (evaluation B C).obj (F.obj a) := + Iso.refl _ + +/-- Whiskering by `F` and then evaluating at `a` is the same as evaluating at `F.obj a`. -/ +@[simp] +theorem whiskeringLeft_comp_evaluation (F : A ⥤ B) (a : A) : + (whiskeringLeft A B C).obj F ⋙ (evaluation A C).obj a = (evaluation B C).obj (F.obj a) := + rfl + +/-- Whiskering by `F` and then evaluating at `a` is the same as evaluating at `F` and then +applying `F`. -/ +@[simps!] +def whiskeringRightCompEvaluation (F : B ⥤ C) (a : A) : + (whiskeringRight A B C).obj F ⋙ (evaluation _ _).obj a ≅ (evaluation _ _).obj a ⋙ F := + Iso.refl _ + +/-- Whiskering by `F` and then evaluating at `a` is the same as evaluating at `F` and then +applying `F`. -/ +@[simp] +theorem whiskeringRight_comp_evaluation (F : B ⥤ C) (a : A) : + (whiskeringRight A B C).obj F ⋙ (evaluation _ _).obj a = (evaluation _ _).obj a ⋙ F := + rfl + variable (A B C) /-- The forward direction for `functorProdFunctorEquiv` -/ From b420076afe5b4ba95e067bfc9add3f4b20b3ba90 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 27 Oct 2024 19:30:22 +0000 Subject: [PATCH 483/505] chore(MeasureTheory/Constructions/Prod/Basic: split (#18286) There were two things going on in this file: * Measurability according to the product sigma-algebra * The product measure The former can be defined much earlier than the latter, in particular without importing `Measure`. --- Archive/Wiedijk100Theorems/BuffonsNeedle.lean | 6 +- Mathlib.lean | 5 +- Mathlib/Analysis/Convolution.lean | 2 +- .../Analysis/Fourier/FourierTransform.lean | 6 +- .../Constructions/HaarToSphere.lean | 2 +- Mathlib/MeasureTheory/Constructions/Pi.lean | 2 +- Mathlib/MeasureTheory/Group/Convolution.lean | 2 +- Mathlib/MeasureTheory/Group/Measure.lean | 2 +- Mathlib/MeasureTheory/Group/Prod.lean | 2 +- .../Integral/DivergenceTheorem.lean | 4 +- Mathlib/MeasureTheory/Integral/Pi.lean | 2 +- .../Prod/Integral.lean => Integral/Prod.lean} | 2 +- .../MeasureTheory/Integral/TorusIntegral.lean | 2 +- .../MeasureTheory/MeasurableSpace/Basic.lean | 7 + .../MeasureTheory/MeasurableSpace/Prod.lean | 136 +++++++++++++++++ .../Measure/FiniteMeasureProd.lean | 2 +- .../MeasureTheory/Measure/Haar/Unique.lean | 8 +- .../Prod/Basic.lean => Measure/Prod.lean} | 139 +----------------- Mathlib/MeasureTheory/PiSystem.lean | 39 +++-- 19 files changed, 194 insertions(+), 176 deletions(-) rename Mathlib/MeasureTheory/{Constructions/Prod/Integral.lean => Integral/Prod.lean} (99%) create mode 100644 Mathlib/MeasureTheory/MeasurableSpace/Prod.lean rename Mathlib/MeasureTheory/{Constructions/Prod/Basic.lean => Measure/Prod.lean} (87%) diff --git a/Archive/Wiedijk100Theorems/BuffonsNeedle.lean b/Archive/Wiedijk100Theorems/BuffonsNeedle.lean index e27b8be9851f2..eff8809144258 100644 --- a/Archive/Wiedijk100Theorems/BuffonsNeedle.lean +++ b/Archive/Wiedijk100Theorems/BuffonsNeedle.lean @@ -4,11 +4,11 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Enrico Z. Borba -/ -import Mathlib.Probability.Density -import Mathlib.Probability.Notation -import Mathlib.MeasureTheory.Constructions.Prod.Integral import Mathlib.Analysis.SpecialFunctions.Integrals +import Mathlib.MeasureTheory.Integral.Prod +import Mathlib.Probability.Density import Mathlib.Probability.Distributions.Uniform +import Mathlib.Probability.Notation /-! diff --git a/Mathlib.lean b/Mathlib.lean index f176ac0021cf7..b6e8beed2f35a 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3324,8 +3324,6 @@ import Mathlib.MeasureTheory.Constructions.HaarToSphere import Mathlib.MeasureTheory.Constructions.Pi import Mathlib.MeasureTheory.Constructions.Polish.Basic import Mathlib.MeasureTheory.Constructions.Polish.EmbeddingReal -import Mathlib.MeasureTheory.Constructions.Prod.Basic -import Mathlib.MeasureTheory.Constructions.Prod.Integral import Mathlib.MeasureTheory.Constructions.Projective import Mathlib.MeasureTheory.Constructions.SubmoduleQuotient import Mathlib.MeasureTheory.Constructions.UnitInterval @@ -3427,6 +3425,7 @@ import Mathlib.MeasureTheory.Integral.MeanInequalities import Mathlib.MeasureTheory.Integral.PeakFunction import Mathlib.MeasureTheory.Integral.Periodic import Mathlib.MeasureTheory.Integral.Pi +import Mathlib.MeasureTheory.Integral.Prod import Mathlib.MeasureTheory.Integral.RieszMarkovKakutani import Mathlib.MeasureTheory.Integral.SetIntegral import Mathlib.MeasureTheory.Integral.SetToL1 @@ -3441,6 +3440,7 @@ import Mathlib.MeasureTheory.MeasurableSpace.Instances import Mathlib.MeasureTheory.MeasurableSpace.Invariants import Mathlib.MeasureTheory.MeasurableSpace.NCard import Mathlib.MeasureTheory.MeasurableSpace.PreorderRestrict +import Mathlib.MeasureTheory.MeasurableSpace.Prod import Mathlib.MeasureTheory.Measure.AEDisjoint import Mathlib.MeasureTheory.Measure.AEMeasurable import Mathlib.MeasureTheory.Measure.AddContent @@ -3478,6 +3478,7 @@ import Mathlib.MeasureTheory.Measure.NullMeasurable import Mathlib.MeasureTheory.Measure.OpenPos import Mathlib.MeasureTheory.Measure.Portmanteau import Mathlib.MeasureTheory.Measure.ProbabilityMeasure +import Mathlib.MeasureTheory.Measure.Prod import Mathlib.MeasureTheory.Measure.Regular import Mathlib.MeasureTheory.Measure.Restrict import Mathlib.MeasureTheory.Measure.SeparableMeasure diff --git a/Mathlib/Analysis/Convolution.lean b/Mathlib/Analysis/Convolution.lean index 00f566fb58ae2..cd64fca205ba0 100644 --- a/Mathlib/Analysis/Convolution.lean +++ b/Mathlib/Analysis/Convolution.lean @@ -5,7 +5,7 @@ Authors: Floris van Doorn -/ import Mathlib.Analysis.Calculus.ContDiff.Basic import Mathlib.Analysis.Calculus.ParametricIntegral -import Mathlib.MeasureTheory.Constructions.Prod.Integral +import Mathlib.MeasureTheory.Integral.Prod import Mathlib.MeasureTheory.Function.LocallyIntegrable import Mathlib.MeasureTheory.Group.Integral import Mathlib.MeasureTheory.Group.Prod diff --git a/Mathlib/Analysis/Fourier/FourierTransform.lean b/Mathlib/Analysis/Fourier/FourierTransform.lean index 8a1b83c8d5e3e..a5d1ddb0067c9 100644 --- a/Mathlib/Analysis/Fourier/FourierTransform.lean +++ b/Mathlib/Analysis/Fourier/FourierTransform.lean @@ -3,13 +3,13 @@ Copyright (c) 2023 David Loeffler. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: David Loeffler -/ +import Mathlib.Algebra.Group.AddChar import Mathlib.Analysis.Complex.Circle import Mathlib.MeasureTheory.Group.Integral +import Mathlib.MeasureTheory.Integral.Prod import Mathlib.MeasureTheory.Integral.SetIntegral -import Mathlib.MeasureTheory.Measure.Haar.OfBasis -import Mathlib.MeasureTheory.Constructions.Prod.Integral import Mathlib.MeasureTheory.Measure.Haar.InnerProductSpace -import Mathlib.Algebra.Group.AddChar +import Mathlib.MeasureTheory.Measure.Haar.OfBasis /-! # The Fourier transform diff --git a/Mathlib/MeasureTheory/Constructions/HaarToSphere.lean b/Mathlib/MeasureTheory/Constructions/HaarToSphere.lean index c236710f3f0fa..8b8f004c2cbd6 100644 --- a/Mathlib/MeasureTheory/Constructions/HaarToSphere.lean +++ b/Mathlib/MeasureTheory/Constructions/HaarToSphere.lean @@ -6,7 +6,7 @@ Authors: Yury Kudryashov import Mathlib.Algebra.Order.Field.Pointwise import Mathlib.Analysis.NormedSpace.SphereNormEquiv import Mathlib.Analysis.SpecialFunctions.Integrals -import Mathlib.MeasureTheory.Constructions.Prod.Integral +import Mathlib.MeasureTheory.Integral.Prod import Mathlib.MeasureTheory.Measure.Lebesgue.EqHaar /-! diff --git a/Mathlib/MeasureTheory/Constructions/Pi.lean b/Mathlib/MeasureTheory/Constructions/Pi.lean index 2a10e3da58fe6..222b04a9c5ac8 100644 --- a/Mathlib/MeasureTheory/Constructions/Pi.lean +++ b/Mathlib/MeasureTheory/Constructions/Pi.lean @@ -3,8 +3,8 @@ Copyright (c) 2020 Floris van Doorn. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn -/ -import Mathlib.MeasureTheory.Constructions.Prod.Basic import Mathlib.MeasureTheory.Group.Measure +import Mathlib.MeasureTheory.Measure.Prod import Mathlib.Topology.Constructions /-! diff --git a/Mathlib/MeasureTheory/Group/Convolution.lean b/Mathlib/MeasureTheory/Group/Convolution.lean index 0796630b4e41a..07a0fe4cdab79 100644 --- a/Mathlib/MeasureTheory/Group/Convolution.lean +++ b/Mathlib/MeasureTheory/Group/Convolution.lean @@ -3,8 +3,8 @@ Copyright (c) 2023 Josha Dekker. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Josha Dekker -/ -import Mathlib.MeasureTheory.Constructions.Prod.Basic import Mathlib.MeasureTheory.Measure.MeasureSpace +import Mathlib.MeasureTheory.Measure.Prod /-! # The multiplicative and additive convolution of measures diff --git a/Mathlib/MeasureTheory/Group/Measure.lean b/Mathlib/MeasureTheory/Group/Measure.lean index 49764e92934c6..c4565fc448a20 100644 --- a/Mathlib/MeasureTheory/Group/Measure.lean +++ b/Mathlib/MeasureTheory/Group/Measure.lean @@ -3,8 +3,8 @@ Copyright (c) 2020 Floris van Doorn. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn -/ -import Mathlib.MeasureTheory.Constructions.Prod.Basic import Mathlib.MeasureTheory.Group.Action +import Mathlib.MeasureTheory.Measure.Prod import Mathlib.Topology.ContinuousMap.CocompactMap /-! diff --git a/Mathlib/MeasureTheory/Group/Prod.lean b/Mathlib/MeasureTheory/Group/Prod.lean index 29e7a7bd7f934..8c7b2f7d77ef1 100644 --- a/Mathlib/MeasureTheory/Group/Prod.lean +++ b/Mathlib/MeasureTheory/Group/Prod.lean @@ -3,8 +3,8 @@ Copyright (c) 2021 Floris van Doorn. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn -/ -import Mathlib.MeasureTheory.Constructions.Prod.Basic import Mathlib.MeasureTheory.Group.Measure +import Mathlib.MeasureTheory.Measure.Prod /-! # Measure theory in the product of groups diff --git a/Mathlib/MeasureTheory/Integral/DivergenceTheorem.lean b/Mathlib/MeasureTheory/Integral/DivergenceTheorem.lean index c42a397ec5d54..827108c784d82 100644 --- a/Mathlib/MeasureTheory/Integral/DivergenceTheorem.lean +++ b/Mathlib/MeasureTheory/Integral/DivergenceTheorem.lean @@ -6,9 +6,9 @@ Authors: Yury Kudryashov import Mathlib.Analysis.BoxIntegral.DivergenceTheorem import Mathlib.Analysis.BoxIntegral.Integrability import Mathlib.Analysis.Calculus.Deriv.Basic -import Mathlib.MeasureTheory.Constructions.Prod.Integral -import Mathlib.MeasureTheory.Integral.IntervalIntegral import Mathlib.Analysis.Calculus.FDeriv.Equiv +import Mathlib.MeasureTheory.Integral.Prod +import Mathlib.MeasureTheory.Integral.IntervalIntegral /-! # Divergence theorem for Bochner integral diff --git a/Mathlib/MeasureTheory/Integral/Pi.lean b/Mathlib/MeasureTheory/Integral/Pi.lean index 3445cc3f930f1..2bf9b9af0c309 100644 --- a/Mathlib/MeasureTheory/Integral/Pi.lean +++ b/Mathlib/MeasureTheory/Integral/Pi.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Xavier Roblot -/ import Mathlib.MeasureTheory.Constructions.Pi -import Mathlib.MeasureTheory.Constructions.Prod.Integral +import Mathlib.MeasureTheory.Integral.Prod /-! # Integration with respect to a finite product of measures diff --git a/Mathlib/MeasureTheory/Constructions/Prod/Integral.lean b/Mathlib/MeasureTheory/Integral/Prod.lean similarity index 99% rename from Mathlib/MeasureTheory/Constructions/Prod/Integral.lean rename to Mathlib/MeasureTheory/Integral/Prod.lean index c86c172d3f6b7..f2ec386e8dfd7 100644 --- a/Mathlib/MeasureTheory/Constructions/Prod/Integral.lean +++ b/Mathlib/MeasureTheory/Integral/Prod.lean @@ -3,9 +3,9 @@ Copyright (c) 2020 Floris van Doorn. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn -/ -import Mathlib.MeasureTheory.Constructions.Prod.Basic import Mathlib.MeasureTheory.Integral.DominatedConvergence import Mathlib.MeasureTheory.Integral.SetIntegral +import Mathlib.MeasureTheory.Measure.Prod /-! # Integration with respect to the product measure diff --git a/Mathlib/MeasureTheory/Integral/TorusIntegral.lean b/Mathlib/MeasureTheory/Integral/TorusIntegral.lean index 9591233916874..80b6f7583b816 100644 --- a/Mathlib/MeasureTheory/Integral/TorusIntegral.lean +++ b/Mathlib/MeasureTheory/Integral/TorusIntegral.lean @@ -3,8 +3,8 @@ Copyright (c) 2022 Cuma Kökmen. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Cuma Kökmen, Yury Kudryashov -/ -import Mathlib.MeasureTheory.Constructions.Prod.Integral import Mathlib.MeasureTheory.Integral.CircleIntegral +import Mathlib.MeasureTheory.Integral.Prod import Mathlib.Order.Fin.Tuple /-! diff --git a/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean b/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean index fa14f94db23e7..2416e12bc0baf 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean @@ -1268,6 +1268,13 @@ theorem isCountablySpanning_measurableSet [MeasurableSpace α] : IsCountablySpanning { s : Set α | MeasurableSet s } := ⟨fun _ => univ, fun _ => MeasurableSet.univ, iUnion_const _⟩ +/-- Rectangles of countably spanning sets are countably spanning. -/ +lemma IsCountablySpanning.prod {C : Set (Set α)} {D : Set (Set β)} (hC : IsCountablySpanning C) + (hD : IsCountablySpanning D) : IsCountablySpanning (image2 (· ×ˢ ·) C D) := by + rcases hC, hD with ⟨⟨s, h1s, h2s⟩, t, h1t, h2t⟩ + refine ⟨fun n => s n.unpair.1 ×ˢ t n.unpair.2, fun n => mem_image2_of_mem (h1s _) (h1t _), ?_⟩ + rw [iUnion_unpair_prod, h2s, h2t, univ_prod_univ] + namespace MeasurableSet /-! diff --git a/Mathlib/MeasureTheory/MeasurableSpace/Prod.lean b/Mathlib/MeasureTheory/MeasurableSpace/Prod.lean new file mode 100644 index 0000000000000..e9f2db14bfc25 --- /dev/null +++ b/Mathlib/MeasureTheory/MeasurableSpace/Prod.lean @@ -0,0 +1,136 @@ +/- +Copyright (c) 2020 Floris van Doorn. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Floris van Doorn +-/ +import Mathlib.MeasureTheory.MeasurableSpace.Embedding +import Mathlib.MeasureTheory.PiSystem + +/-! +# The product sigma algebra + +This file talks about the measurability of operations on binary functions. +-/ + +assert_not_exists MeasureTheory.Measure + +noncomputable section + +open Set MeasurableSpace + +variable {α β γ : Type*} [MeasurableSpace α] [MeasurableSpace β] [MeasurableSpace γ] + +/-- The product of generated σ-algebras is the one generated by rectangles, if both generating sets + are countably spanning. -/ +theorem generateFrom_prod_eq {α β} {C : Set (Set α)} {D : Set (Set β)} (hC : IsCountablySpanning C) + (hD : IsCountablySpanning D) : + @Prod.instMeasurableSpace _ _ (generateFrom C) (generateFrom D) = + generateFrom (image2 (· ×ˢ ·) C D) := by + apply le_antisymm + · refine sup_le ?_ ?_ <;> rw [comap_generateFrom] <;> apply generateFrom_le <;> + rintro _ ⟨s, hs, rfl⟩ + · rcases hD with ⟨t, h1t, h2t⟩ + rw [← prod_univ, ← h2t, prod_iUnion] + apply MeasurableSet.iUnion + intro n + apply measurableSet_generateFrom + exact ⟨s, hs, t n, h1t n, rfl⟩ + · rcases hC with ⟨t, h1t, h2t⟩ + rw [← univ_prod, ← h2t, iUnion_prod_const] + apply MeasurableSet.iUnion + rintro n + apply measurableSet_generateFrom + exact mem_image2_of_mem (h1t n) hs + · apply generateFrom_le + rintro _ ⟨s, hs, t, ht, rfl⟩ + dsimp only + rw [prod_eq] + apply (measurable_fst _).inter (measurable_snd _) + · exact measurableSet_generateFrom hs + · exact measurableSet_generateFrom ht + +/-- If `C` and `D` generate the σ-algebras on `α` resp. `β`, then rectangles formed by `C` and `D` + generate the σ-algebra on `α × β`. -/ +theorem generateFrom_eq_prod {C : Set (Set α)} {D : Set (Set β)} (hC : generateFrom C = ‹_›) + (hD : generateFrom D = ‹_›) (h2C : IsCountablySpanning C) (h2D : IsCountablySpanning D) : + generateFrom (image2 (· ×ˢ ·) C D) = Prod.instMeasurableSpace := by + rw [← hC, ← hD, generateFrom_prod_eq h2C h2D] + +/-- The product σ-algebra is generated from boxes, i.e. `s ×ˢ t` for sets `s : Set α` and + `t : Set β`. -/ +lemma generateFrom_prod : + generateFrom (image2 (· ×ˢ ·) { s : Set α | MeasurableSet s } { t : Set β | MeasurableSet t }) = + Prod.instMeasurableSpace := + generateFrom_eq_prod generateFrom_measurableSet generateFrom_measurableSet + isCountablySpanning_measurableSet isCountablySpanning_measurableSet + +/-- Rectangles form a π-system. -/ +lemma isPiSystem_prod : + IsPiSystem (image2 (· ×ˢ ·) { s : Set α | MeasurableSet s } { t : Set β | MeasurableSet t }) := + isPiSystem_measurableSet.prod isPiSystem_measurableSet + +lemma MeasurableEmbedding.prod_mk {α β γ δ : Type*} {mα : MeasurableSpace α} + {mβ : MeasurableSpace β} {mγ : MeasurableSpace γ} {mδ : MeasurableSpace δ} {f : α → β} + {g : γ → δ} (hg : MeasurableEmbedding g) (hf : MeasurableEmbedding f) : + MeasurableEmbedding fun x : γ × α => (g x.1, f x.2) := by + have h_inj : Function.Injective fun x : γ × α => (g x.fst, f x.snd) := by + intro x y hxy + rw [← @Prod.mk.eta _ _ x, ← @Prod.mk.eta _ _ y] + simp only [Prod.mk.inj_iff] at hxy ⊢ + exact ⟨hg.injective hxy.1, hf.injective hxy.2⟩ + refine ⟨h_inj, ?_, ?_⟩ + · exact (hg.measurable.comp measurable_fst).prod_mk (hf.measurable.comp measurable_snd) + · -- Induction using the π-system of rectangles + refine fun s hs => + @MeasurableSpace.induction_on_inter _ + (fun s => MeasurableSet ((fun x : γ × α => (g x.fst, f x.snd)) '' s)) _ _ + generateFrom_prod.symm isPiSystem_prod ?_ ?_ ?_ ?_ _ hs + · simp only [Set.image_empty, MeasurableSet.empty] + · rintro t ⟨t₁, ht₁, t₂, ht₂, rfl⟩ + rw [← Set.prod_image_image_eq] + exact (hg.measurableSet_image.mpr ht₁).prod (hf.measurableSet_image.mpr ht₂) + · intro t _ ht_m + rw [← Set.range_diff_image h_inj, ← Set.prod_range_range_eq] + exact + MeasurableSet.diff (MeasurableSet.prod hg.measurableSet_range hf.measurableSet_range) ht_m + · intro g _ _ hg + simp_rw [Set.image_iUnion] + exact MeasurableSet.iUnion hg + +lemma MeasurableEmbedding.prod_mk_left {β γ : Type*} [MeasurableSingletonClass α] + {mβ : MeasurableSpace β} {mγ : MeasurableSpace γ} + (x : α) {f : γ → β} (hf : MeasurableEmbedding f) : + MeasurableEmbedding (fun y ↦ (x, f y)) where + injective := by + intro y y' + simp only [Prod.mk.injEq, true_and] + exact fun h ↦ hf.injective h + measurable := Measurable.prod_mk measurable_const hf.measurable + measurableSet_image' := by + intro s hs + convert (MeasurableSet.singleton x).prod (hf.measurableSet_image.mpr hs) + ext x + simp + +lemma measurableEmbedding_prod_mk_left [MeasurableSingletonClass α] (x : α) : + MeasurableEmbedding (Prod.mk x : β → α × β) := + MeasurableEmbedding.prod_mk_left x MeasurableEmbedding.id + +lemma MeasurableEmbedding.prod_mk_right {β γ : Type*} [MeasurableSingletonClass α] + {mβ : MeasurableSpace β} {mγ : MeasurableSpace γ} + {f : γ → β} (hf : MeasurableEmbedding f) (x : α) : + MeasurableEmbedding (fun y ↦ (f y, x)) where + injective := by + intro y y' + simp only [Prod.mk.injEq, and_true] + exact fun h ↦ hf.injective h + measurable := Measurable.prod_mk hf.measurable measurable_const + measurableSet_image' := by + intro s hs + convert (hf.measurableSet_image.mpr hs).prod (MeasurableSet.singleton x) + ext x + simp + +lemma measurableEmbedding_prod_mk_right [MeasurableSingletonClass α] (x : α) : + MeasurableEmbedding (fun y ↦ (y, x) : β → β × α) := + MeasurableEmbedding.prod_mk_right MeasurableEmbedding.id x diff --git a/Mathlib/MeasureTheory/Measure/FiniteMeasureProd.lean b/Mathlib/MeasureTheory/Measure/FiniteMeasureProd.lean index 97ffeece7c11c..422da03a98cea 100644 --- a/Mathlib/MeasureTheory/Measure/FiniteMeasureProd.lean +++ b/Mathlib/MeasureTheory/Measure/FiniteMeasureProd.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kalle Kytölä -/ import Mathlib.MeasureTheory.Measure.ProbabilityMeasure -import Mathlib.MeasureTheory.Constructions.Prod.Basic +import Mathlib.MeasureTheory.Measure.Prod /-! # Products of finite measures and probability measures diff --git a/Mathlib/MeasureTheory/Measure/Haar/Unique.lean b/Mathlib/MeasureTheory/Measure/Haar/Unique.lean index 0b75436b65433..a9088acc3aeda 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/Unique.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/Unique.lean @@ -3,14 +3,14 @@ Copyright (c) 2023 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ -import Mathlib.MeasureTheory.Constructions.Prod.Integral import Mathlib.MeasureTheory.Function.LocallyIntegrable import Mathlib.MeasureTheory.Group.Integral +import Mathlib.MeasureTheory.Integral.Prod +import Mathlib.MeasureTheory.Integral.SetIntegral +import Mathlib.MeasureTheory.Measure.EverywherePos +import Mathlib.MeasureTheory.Measure.Haar.Basic import Mathlib.Topology.Metrizable.Urysohn import Mathlib.Topology.UrysohnsLemma -import Mathlib.MeasureTheory.Measure.Haar.Basic -import Mathlib.MeasureTheory.Measure.EverywherePos -import Mathlib.MeasureTheory.Integral.SetIntegral /-! # Uniqueness of Haar measure in locally compact groups diff --git a/Mathlib/MeasureTheory/Constructions/Prod/Basic.lean b/Mathlib/MeasureTheory/Measure/Prod.lean similarity index 87% rename from Mathlib/MeasureTheory/Constructions/Prod/Basic.lean rename to Mathlib/MeasureTheory/Measure/Prod.lean index 7b687ed495686..76a9fe13f08da 100644 --- a/Mathlib/MeasureTheory/Constructions/Prod/Basic.lean +++ b/Mathlib/MeasureTheory/Measure/Prod.lean @@ -3,9 +3,10 @@ Copyright (c) 2020 Floris van Doorn. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn -/ -import Mathlib.MeasureTheory.Measure.GiryMonad import Mathlib.Dynamics.Ergodic.MeasurePreserving import Mathlib.MeasureTheory.Integral.Lebesgue +import Mathlib.MeasureTheory.MeasurableSpace.Prod +import Mathlib.MeasureTheory.Measure.GiryMonad import Mathlib.MeasureTheory.Measure.OpenPos /-! @@ -62,81 +63,11 @@ open Filter hiding prod_eq map variable {α α' β β' γ E : Type*} -/-- Rectangles formed by π-systems form a π-system. -/ -theorem IsPiSystem.prod {C : Set (Set α)} {D : Set (Set β)} (hC : IsPiSystem C) - (hD : IsPiSystem D) : IsPiSystem (image2 (· ×ˢ ·) C D) := by - rintro _ ⟨s₁, hs₁, t₁, ht₁, rfl⟩ _ ⟨s₂, hs₂, t₂, ht₂, rfl⟩ hst - rw [prod_inter_prod] at hst ⊢; rw [prod_nonempty_iff] at hst - exact mem_image2_of_mem (hC _ hs₁ _ hs₂ hst.1) (hD _ ht₁ _ ht₂ hst.2) - -/-- Rectangles of countably spanning sets are countably spanning. -/ -theorem IsCountablySpanning.prod {C : Set (Set α)} {D : Set (Set β)} (hC : IsCountablySpanning C) - (hD : IsCountablySpanning D) : IsCountablySpanning (image2 (· ×ˢ ·) C D) := by - rcases hC, hD with ⟨⟨s, h1s, h2s⟩, t, h1t, h2t⟩ - refine ⟨fun n => s n.unpair.1 ×ˢ t n.unpair.2, fun n => mem_image2_of_mem (h1s _) (h1t _), ?_⟩ - rw [iUnion_unpair_prod, h2s, h2t, univ_prod_univ] - variable [MeasurableSpace α] [MeasurableSpace α'] [MeasurableSpace β] [MeasurableSpace β'] variable [MeasurableSpace γ] variable {μ μ' : Measure α} {ν ν' : Measure β} {τ : Measure γ} variable [NormedAddCommGroup E] -/-! ### Measurability - -Before we define the product measure, we can talk about the measurability of operations on binary -functions. We show that if `f` is a binary measurable function, then the function that integrates -along one of the variables (using either the Lebesgue or Bochner integral) is measurable. --/ - -/-- The product of generated σ-algebras is the one generated by rectangles, if both generating sets - are countably spanning. -/ -theorem generateFrom_prod_eq {α β} {C : Set (Set α)} {D : Set (Set β)} (hC : IsCountablySpanning C) - (hD : IsCountablySpanning D) : - @Prod.instMeasurableSpace _ _ (generateFrom C) (generateFrom D) = - generateFrom (image2 (· ×ˢ ·) C D) := by - apply le_antisymm - · refine sup_le ?_ ?_ <;> rw [comap_generateFrom] <;> apply generateFrom_le <;> - rintro _ ⟨s, hs, rfl⟩ - · rcases hD with ⟨t, h1t, h2t⟩ - rw [← prod_univ, ← h2t, prod_iUnion] - apply MeasurableSet.iUnion - intro n - apply measurableSet_generateFrom - exact ⟨s, hs, t n, h1t n, rfl⟩ - · rcases hC with ⟨t, h1t, h2t⟩ - rw [← univ_prod, ← h2t, iUnion_prod_const] - apply MeasurableSet.iUnion - rintro n - apply measurableSet_generateFrom - exact mem_image2_of_mem (h1t n) hs - · apply generateFrom_le - rintro _ ⟨s, hs, t, ht, rfl⟩ - dsimp only - rw [prod_eq] - apply (measurable_fst _).inter (measurable_snd _) - · exact measurableSet_generateFrom hs - · exact measurableSet_generateFrom ht - -/-- If `C` and `D` generate the σ-algebras on `α` resp. `β`, then rectangles formed by `C` and `D` - generate the σ-algebra on `α × β`. -/ -theorem generateFrom_eq_prod {C : Set (Set α)} {D : Set (Set β)} (hC : generateFrom C = ‹_›) - (hD : generateFrom D = ‹_›) (h2C : IsCountablySpanning C) (h2D : IsCountablySpanning D) : - generateFrom (image2 (· ×ˢ ·) C D) = Prod.instMeasurableSpace := by - rw [← hC, ← hD, generateFrom_prod_eq h2C h2D] - -/-- The product σ-algebra is generated from boxes, i.e. `s ×ˢ t` for sets `s : Set α` and - `t : Set β`. -/ -theorem generateFrom_prod : - generateFrom (image2 (· ×ˢ ·) { s : Set α | MeasurableSet s } { t : Set β | MeasurableSet t }) = - Prod.instMeasurableSpace := - generateFrom_eq_prod generateFrom_measurableSet generateFrom_measurableSet - isCountablySpanning_measurableSet isCountablySpanning_measurableSet - -/-- Rectangles form a π-system. -/ -theorem isPiSystem_prod : - IsPiSystem (image2 (· ×ˢ ·) { s : Set α | MeasurableSet s } { t : Set β | MeasurableSet t }) := - isPiSystem_measurableSet.prod isPiSystem_measurableSet - /-- If `ν` is a finite measure, and `s ⊆ α × β` is measurable, then `x ↦ ν { y | (x, y) ∈ s }` is a measurable function. `measurable_measure_prod_mk_left` is strictly more general. -/ theorem measurable_measure_prod_mk_left_finite [IsFiniteMeasure ν] {s : Set (α × β)} @@ -185,72 +116,6 @@ theorem Measurable.map_prod_mk_right {μ : Measure α} [SFinite μ] : simp_rw [map_apply measurable_prod_mk_right hs] exact measurable_measure_prod_mk_right hs -theorem MeasurableEmbedding.prod_mk {α β γ δ : Type*} {mα : MeasurableSpace α} - {mβ : MeasurableSpace β} {mγ : MeasurableSpace γ} {mδ : MeasurableSpace δ} {f : α → β} - {g : γ → δ} (hg : MeasurableEmbedding g) (hf : MeasurableEmbedding f) : - MeasurableEmbedding fun x : γ × α => (g x.1, f x.2) := by - have h_inj : Function.Injective fun x : γ × α => (g x.fst, f x.snd) := by - intro x y hxy - rw [← @Prod.mk.eta _ _ x, ← @Prod.mk.eta _ _ y] - simp only [Prod.mk.inj_iff] at hxy ⊢ - exact ⟨hg.injective hxy.1, hf.injective hxy.2⟩ - refine ⟨h_inj, ?_, ?_⟩ - · exact (hg.measurable.comp measurable_fst).prod_mk (hf.measurable.comp measurable_snd) - · -- Induction using the π-system of rectangles - refine fun s hs => - @MeasurableSpace.induction_on_inter _ - (fun s => MeasurableSet ((fun x : γ × α => (g x.fst, f x.snd)) '' s)) _ _ - generateFrom_prod.symm isPiSystem_prod ?_ ?_ ?_ ?_ _ hs - · simp only [Set.image_empty, MeasurableSet.empty] - · rintro t ⟨t₁, ht₁, t₂, ht₂, rfl⟩ - rw [← Set.prod_image_image_eq] - exact (hg.measurableSet_image.mpr ht₁).prod (hf.measurableSet_image.mpr ht₂) - · intro t _ ht_m - rw [← Set.range_diff_image h_inj, ← Set.prod_range_range_eq] - exact - MeasurableSet.diff (MeasurableSet.prod hg.measurableSet_range hf.measurableSet_range) ht_m - · intro g _ _ hg - simp_rw [Set.image_iUnion] - exact MeasurableSet.iUnion hg - -lemma MeasurableEmbedding.prod_mk_left {β γ : Type*} [MeasurableSingletonClass α] - {mβ : MeasurableSpace β} {mγ : MeasurableSpace γ} - (x : α) {f : γ → β} (hf : MeasurableEmbedding f) : - MeasurableEmbedding (fun y ↦ (x, f y)) where - injective := by - intro y y' - simp only [Prod.mk.injEq, true_and] - exact fun h ↦ hf.injective h - measurable := Measurable.prod_mk measurable_const hf.measurable - measurableSet_image' := by - intro s hs - convert (MeasurableSet.singleton x).prod (hf.measurableSet_image.mpr hs) - ext x - simp - -lemma measurableEmbedding_prod_mk_left [MeasurableSingletonClass α] (x : α) : - MeasurableEmbedding (Prod.mk x : β → α × β) := - MeasurableEmbedding.prod_mk_left x MeasurableEmbedding.id - -lemma MeasurableEmbedding.prod_mk_right {β γ : Type*} [MeasurableSingletonClass α] - {mβ : MeasurableSpace β} {mγ : MeasurableSpace γ} - {f : γ → β} (hf : MeasurableEmbedding f) (x : α) : - MeasurableEmbedding (fun y ↦ (f y, x)) where - injective := by - intro y y' - simp only [Prod.mk.injEq, and_true] - exact fun h ↦ hf.injective h - measurable := Measurable.prod_mk hf.measurable measurable_const - measurableSet_image' := by - intro s hs - convert (hf.measurableSet_image.mpr hs).prod (MeasurableSet.singleton x) - ext x - simp - -lemma measurableEmbedding_prod_mk_right [MeasurableSingletonClass α] (x : α) : - MeasurableEmbedding (fun y ↦ (y, x) : β → β × α) := - MeasurableEmbedding.prod_mk_right MeasurableEmbedding.id x - /-- The Lebesgue integral is measurable. This shows that the integrand of (the right-hand-side of) Tonelli's theorem is measurable. -/ theorem Measurable.lintegral_prod_right' [SFinite ν] : diff --git a/Mathlib/MeasureTheory/PiSystem.lean b/Mathlib/MeasureTheory/PiSystem.lean index fdcbf376f1d77..2f16137420ef0 100644 --- a/Mathlib/MeasureTheory/PiSystem.lean +++ b/Mathlib/MeasureTheory/PiSystem.lean @@ -58,10 +58,12 @@ open MeasurableSpace Set open MeasureTheory +variable {α β : Type*} + /-- A π-system is a collection of subsets of `α` that is closed under binary intersection of non-disjoint sets. Usually it is also required that the collection is nonempty, but we don't do that here. -/ -def IsPiSystem {α} (C : Set (Set α)) : Prop := +def IsPiSystem (C : Set (Set α)) : Prop := ∀ᵉ (s ∈ C) (t ∈ C), (s ∩ t : Set α).Nonempty → s ∩ t ∈ C namespace MeasurableSpace @@ -71,12 +73,12 @@ theorem isPiSystem_measurableSet {α : Type*} [MeasurableSpace α] : end MeasurableSpace -theorem IsPiSystem.singleton {α} (S : Set α) : IsPiSystem ({S} : Set (Set α)) := by +theorem IsPiSystem.singleton (S : Set α) : IsPiSystem ({S} : Set (Set α)) := by intro s h_s t h_t _ rw [Set.mem_singleton_iff.1 h_s, Set.mem_singleton_iff.1 h_t, Set.inter_self, Set.mem_singleton_iff] -theorem IsPiSystem.insert_empty {α} {S : Set (Set α)} (h_pi : IsPiSystem S) : +theorem IsPiSystem.insert_empty {S : Set (Set α)} (h_pi : IsPiSystem S) : IsPiSystem (insert ∅ S) := by intro s hs t ht hst cases' hs with hs hs @@ -85,7 +87,7 @@ theorem IsPiSystem.insert_empty {α} {S : Set (Set α)} (h_pi : IsPiSystem S) : · simp [ht] · exact Set.mem_insert_of_mem _ (h_pi s hs t ht hst) -theorem IsPiSystem.insert_univ {α} {S : Set (Set α)} (h_pi : IsPiSystem S) : +theorem IsPiSystem.insert_univ {S : Set (Set α)} (h_pi : IsPiSystem S) : IsPiSystem (insert Set.univ S) := by intro s hs t ht hst cases' hs with hs hs @@ -114,9 +116,16 @@ theorem isPiSystem_iUnion_of_monotone {α ι} [SemilatticeSup ι] (p : ι → Se (hp_pi : ∀ n, IsPiSystem (p n)) (hp_mono : Monotone p) : IsPiSystem (⋃ n, p n) := isPiSystem_iUnion_of_directed_le p hp_pi (Monotone.directed_le hp_mono) +/-- Rectangles formed by π-systems form a π-system. -/ +lemma IsPiSystem.prod {C : Set (Set α)} {D : Set (Set β)} (hC : IsPiSystem C) (hD : IsPiSystem D) : + IsPiSystem (image2 (· ×ˢ ·) C D) := by + rintro _ ⟨s₁, hs₁, t₁, ht₁, rfl⟩ _ ⟨s₂, hs₂, t₂, ht₂, rfl⟩ hst + rw [prod_inter_prod] at hst ⊢; rw [prod_nonempty_iff] at hst + exact mem_image2_of_mem (hC _ hs₁ _ hs₂ hst.1) (hD _ ht₁ _ ht₂ hst.2) + section Order -variable {α : Type*} {ι ι' : Sort*} [LinearOrder α] +variable {ι ι' : Sort*} [LinearOrder α] theorem isPiSystem_image_Iio (s : Set α) : IsPiSystem (Iio '' s) := by rintro _ ⟨a, ha, rfl⟩ _ ⟨b, hb, rfl⟩ - @@ -194,45 +203,45 @@ end Order /-- Given a collection `S` of subsets of `α`, then `generatePiSystem S` is the smallest π-system containing `S`. -/ -inductive generatePiSystem {α} (S : Set (Set α)) : Set (Set α) +inductive generatePiSystem (S : Set (Set α)) : Set (Set α) | base {s : Set α} (h_s : s ∈ S) : generatePiSystem S s | inter {s t : Set α} (h_s : generatePiSystem S s) (h_t : generatePiSystem S t) (h_nonempty : (s ∩ t).Nonempty) : generatePiSystem S (s ∩ t) -theorem isPiSystem_generatePiSystem {α} (S : Set (Set α)) : IsPiSystem (generatePiSystem S) := +theorem isPiSystem_generatePiSystem (S : Set (Set α)) : IsPiSystem (generatePiSystem S) := fun _ h_s _ h_t h_nonempty => generatePiSystem.inter h_s h_t h_nonempty -theorem subset_generatePiSystem_self {α} (S : Set (Set α)) : S ⊆ generatePiSystem S := fun _ => +theorem subset_generatePiSystem_self (S : Set (Set α)) : S ⊆ generatePiSystem S := fun _ => generatePiSystem.base -theorem generatePiSystem_subset_self {α} {S : Set (Set α)} (h_S : IsPiSystem S) : +theorem generatePiSystem_subset_self {S : Set (Set α)} (h_S : IsPiSystem S) : generatePiSystem S ⊆ S := fun x h => by induction h with | base h_s => exact h_s | inter _ _ h_nonempty h_s h_u => exact h_S _ h_s _ h_u h_nonempty -theorem generatePiSystem_eq {α} {S : Set (Set α)} (h_pi : IsPiSystem S) : generatePiSystem S = S := +theorem generatePiSystem_eq {S : Set (Set α)} (h_pi : IsPiSystem S) : generatePiSystem S = S := Set.Subset.antisymm (generatePiSystem_subset_self h_pi) (subset_generatePiSystem_self S) -theorem generatePiSystem_mono {α} {S T : Set (Set α)} (hST : S ⊆ T) : +theorem generatePiSystem_mono {S T : Set (Set α)} (hST : S ⊆ T) : generatePiSystem S ⊆ generatePiSystem T := fun t ht => by induction ht with | base h_s => exact generatePiSystem.base (Set.mem_of_subset_of_mem hST h_s) | inter _ _ h_nonempty h_s h_u => exact isPiSystem_generatePiSystem T _ h_s _ h_u h_nonempty -theorem generatePiSystem_measurableSet {α} [M : MeasurableSpace α] {S : Set (Set α)} +theorem generatePiSystem_measurableSet [M : MeasurableSpace α] {S : Set (Set α)} (h_meas_S : ∀ s ∈ S, MeasurableSet s) (t : Set α) (h_in_pi : t ∈ generatePiSystem S) : MeasurableSet t := by induction h_in_pi with | base h_s => apply h_meas_S _ h_s | inter _ _ _ h_s h_u => apply MeasurableSet.inter h_s h_u -theorem generateFrom_measurableSet_of_generatePiSystem {α} {g : Set (Set α)} (t : Set α) +theorem generateFrom_measurableSet_of_generatePiSystem {g : Set (Set α)} (t : Set α) (ht : t ∈ generatePiSystem g) : MeasurableSet[generateFrom g] t := @generatePiSystem_measurableSet α (generateFrom g) g (fun _ h_s_in_g => measurableSet_generateFrom h_s_in_g) t ht -theorem generateFrom_generatePiSystem_eq {α} {g : Set (Set α)} : +theorem generateFrom_generatePiSystem_eq {g : Set (Set α)} : generateFrom (generatePiSystem g) = generateFrom g := by apply le_antisymm <;> apply generateFrom_le · exact fun t h_t => generateFrom_measurableSet_of_generatePiSystem t h_t @@ -535,7 +544,7 @@ theorem has_diff {s₁ s₂ : Set α} (h₁ : d.Has s₁) (h₂ : d.Has s₂) (h instance instLEDynkinSystem : LE (DynkinSystem α) where le m₁ m₂ := m₁.Has ≤ m₂.Has -theorem le_def {α} {a b : DynkinSystem α} : a ≤ b ↔ a.Has ≤ b.Has := +theorem le_def {a b : DynkinSystem α} : a ≤ b ↔ a.Has ≤ b.Has := Iff.rfl instance : PartialOrder (DynkinSystem α) := From 04f438a04f3eb6279c293f133bdcf717d785331d Mon Sep 17 00:00:00 2001 From: Amelia Livingston <101damnations@github.com> Date: Sun, 27 Oct 2024 19:43:22 +0000 Subject: [PATCH 484/505] feat: tensor products of bialgebras (#11977) We define the data in the monoidal structure on bialgebras - e.g. the bialgebra instance on a tensor product of bialgebras, and the tensor product of two `BialgHom`s as a `BialgHom`. This is done by combining the corresponding API for coalgebras and algebras. Co-authored-by: Amelia Livingston Co-authored-by: Eric Wieser --- Mathlib.lean | 1 + .../RingTheory/Bialgebra/TensorProduct.lean | 180 ++++++++++++++++++ Mathlib/RingTheory/TensorProduct/Basic.lean | 49 +++++ 3 files changed, 230 insertions(+) create mode 100644 Mathlib/RingTheory/Bialgebra/TensorProduct.lean diff --git a/Mathlib.lean b/Mathlib.lean index b6e8beed2f35a..8613416a68ff8 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3975,6 +3975,7 @@ import Mathlib.RingTheory.Bezout import Mathlib.RingTheory.Bialgebra.Basic import Mathlib.RingTheory.Bialgebra.Equiv import Mathlib.RingTheory.Bialgebra.Hom +import Mathlib.RingTheory.Bialgebra.TensorProduct import Mathlib.RingTheory.Binomial import Mathlib.RingTheory.ChainOfDivisors import Mathlib.RingTheory.ClassGroup diff --git a/Mathlib/RingTheory/Bialgebra/TensorProduct.lean b/Mathlib/RingTheory/Bialgebra/TensorProduct.lean new file mode 100644 index 0000000000000..1a461c3252fab --- /dev/null +++ b/Mathlib/RingTheory/Bialgebra/TensorProduct.lean @@ -0,0 +1,180 @@ +/- +Copyright (c) 2024 Amelia Livingston. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Amelia Livingston +-/ +import Mathlib.RingTheory.Bialgebra.Equiv +import Mathlib.RingTheory.Coalgebra.TensorProduct +import Mathlib.RingTheory.TensorProduct.Basic + +/-! +# Tensor products of bialgebras + +We define the data in the monoidal structure on the category of bialgebras - e.g. the bialgebra +instance on a tensor product of bialgebras, and the tensor product of two `BialgHom`s as a +`BialgHom`. This is done by combining the corresponding API for coalgebras and algebras. + +## Implementation notes + +Since the coalgebra instance on a tensor product of coalgebras is defined using a universe +monomorphic monoidal category structure on `ModuleCat R`, we require our bialgebras to be in the +same universe as the base ring `R` here too. Any contribution that achieves universe polymorphism +would be welcome. For instance, the tensor product of bialgebras in the +[FLT repo](https://github.com/ImperialCollegeLondon/FLT/blob/eef74b4538c8852363936dfaad23e6ffba72eca5/FLT/mathlibExperiments/Coalgebra/TensorProduct.lean) +is already universe polymorphic since it does not go via category theory. + +-/ + +universe v u + +open scoped TensorProduct + +namespace Bialgebra.TensorProduct + +variable (R A B : Type*) [CommRing R] [Ring A] [Ring B] [Bialgebra R A] [Bialgebra R B] + +lemma counit_eq_algHom_toLinearMap : + Coalgebra.counit (R := R) (A := A ⊗[R] B) = + ((Algebra.TensorProduct.lmul' R).comp (Algebra.TensorProduct.map + (Bialgebra.counitAlgHom R A) (Bialgebra.counitAlgHom R B))).toLinearMap := by + rfl + +lemma comul_eq_algHom_toLinearMap : + Coalgebra.comul (R := R) (A := A ⊗[R] B) = + ((Algebra.TensorProduct.tensorTensorTensorComm R A A B B).toAlgHom.comp + (Algebra.TensorProduct.map (Bialgebra.comulAlgHom R A) + (Bialgebra.comulAlgHom R B))).toLinearMap := by + rfl + +end Bialgebra.TensorProduct + +noncomputable instance TensorProduct.instBialgebra + {R A B : Type u} [CommRing R] [Ring A] [Ring B] + [Bialgebra R A] [Bialgebra R B] : Bialgebra R (A ⊗[R] B) := by + have hcounit := congr(DFunLike.coe $(Bialgebra.TensorProduct.counit_eq_algHom_toLinearMap R A B)) + have hcomul := congr(DFunLike.coe $(Bialgebra.TensorProduct.comul_eq_algHom_toLinearMap R A B)) + refine Bialgebra.mk' R (A ⊗[R] B) ?_ (fun {x y} => ?_) ?_ (fun {x y} => ?_) <;> + simp_all only [AlgHom.toLinearMap_apply] <;> + simp only [_root_.map_one, _root_.map_mul] + +namespace Bialgebra.TensorProduct + +variable {R A B C D : Type u} [CommRing R] [Ring A] [Ring B] [Ring C] [Ring D] + [Bialgebra R A] [Bialgebra R B] [Bialgebra R C] [Bialgebra R D] + +/-- The tensor product of two bialgebra morphisms as a bialgebra morphism. -/ +noncomputable def map (f : A →ₐc[R] B) (g : C →ₐc[R] D) : + A ⊗[R] C →ₐc[R] B ⊗[R] D := + { Coalgebra.TensorProduct.map (f : A →ₗc[R] B) (g : C →ₗc[R] D), + Algebra.TensorProduct.map (f : A →ₐ[R] B) (g : C →ₐ[R] D) with } + +@[simp] +theorem map_tmul (f : A →ₐc[R] B) (g : C →ₐc[R] D) (x : A) (y : C) : + map f g (x ⊗ₜ y) = f x ⊗ₜ g y := + rfl + +@[simp] +theorem map_toCoalgHom (f : A →ₐc[R] B) (g : C →ₐc[R] D) : + map f g = Coalgebra.TensorProduct.map (f : A →ₗc[R] B) (g : C →ₗc[R] D) := rfl + +@[simp] +theorem map_toAlgHom (f : A →ₐc[R] B) (g : C →ₐc[R] D) : + (map f g : A ⊗[R] C →ₐ[R] B ⊗[R] D) = + Algebra.TensorProduct.map (f : A →ₐ[R] B) (g : C →ₐ[R] D) := + rfl + +variable (R A B C) in +/-- The associator for tensor products of R-bialgebras, as a bialgebra equivalence. -/ +protected noncomputable def assoc : + (A ⊗[R] B) ⊗[R] C ≃ₐc[R] A ⊗[R] (B ⊗[R] C) := + { Coalgebra.TensorProduct.assoc R A B C, Algebra.TensorProduct.assoc R A B C with } + +@[simp] +theorem assoc_tmul (x : A) (y : B) (z : C) : + Bialgebra.TensorProduct.assoc R A B C ((x ⊗ₜ y) ⊗ₜ z) = x ⊗ₜ (y ⊗ₜ z) := + rfl + +@[simp] +theorem assoc_symm_tmul (x : A) (y : B) (z : C) : + (Bialgebra.TensorProduct.assoc R A B C).symm (x ⊗ₜ (y ⊗ₜ z)) = (x ⊗ₜ y) ⊗ₜ z := + rfl + +@[simp] +theorem assoc_toCoalgEquiv : + (Bialgebra.TensorProduct.assoc R A B C : _ ≃ₗc[R] _) = + Coalgebra.TensorProduct.assoc R A B C := rfl + +@[simp] +theorem assoc_toAlgEquiv : + (Bialgebra.TensorProduct.assoc R A B C : _ ≃ₐ[R] _) = + Algebra.TensorProduct.assoc R A B C := rfl + +variable (R A) in +/-- The base ring is a left identity for the tensor product of bialgebras, up to +bialgebra equivalence. -/ +protected noncomputable def lid : R ⊗[R] A ≃ₐc[R] A := + { Coalgebra.TensorProduct.lid R A, Algebra.TensorProduct.lid R A with } + +@[simp] +theorem lid_toCoalgEquiv : + (Bialgebra.TensorProduct.lid R A : R ⊗[R] A ≃ₗc[R] A) = Coalgebra.TensorProduct.lid R A := rfl + +@[simp] +theorem lid_toAlgEquiv : + (Bialgebra.TensorProduct.lid R A : R ⊗[R] A ≃ₐ[R] A) = Algebra.TensorProduct.lid R A := rfl + +@[simp] +theorem lid_tmul (r : R) (a : A) : Bialgebra.TensorProduct.lid R A (r ⊗ₜ a) = r • a := rfl + +@[simp] +theorem lid_symm_apply (a : A) : (Bialgebra.TensorProduct.lid R A).symm a = 1 ⊗ₜ a := rfl + +/- TODO: make this defeq, which would involve adding a heterobasic version of +`Coalgebra.TensorProduct.rid`. -/ +theorem coalgebra_rid_eq_algebra_rid_apply (x : A ⊗[R] R) : + Coalgebra.TensorProduct.rid R A x = Algebra.TensorProduct.rid R R A x := + congr($((TensorProduct.AlgebraTensorModule.rid_eq_rid R A).symm) x) + +variable (R A) in +/-- The base ring is a right identity for the tensor product of bialgebras, up to +bialgebra equivalence. -/ +protected noncomputable def rid : A ⊗[R] R ≃ₐc[R] A where + toCoalgEquiv := Coalgebra.TensorProduct.rid R A + map_mul' x y := by + simp only [CoalgEquiv.toCoalgHom_eq_coe, CoalgHom.toLinearMap_eq_coe, AddHom.toFun_eq_coe, + LinearMap.coe_toAddHom, CoalgHom.coe_toLinearMap, CoalgHom.coe_coe, + coalgebra_rid_eq_algebra_rid_apply, map_mul] + +@[simp] +theorem rid_toCoalgEquiv : + (Bialgebra.TensorProduct.rid R A : A ⊗[R] R ≃ₗc[R] A) = Coalgebra.TensorProduct.rid R A := rfl + +@[simp] +theorem rid_toAlgEquiv : + (Bialgebra.TensorProduct.rid R A : A ⊗[R] R ≃ₐ[R] A) = Algebra.TensorProduct.rid R R A := by + ext x + exact coalgebra_rid_eq_algebra_rid_apply x + +@[simp] +theorem rid_tmul (r : R) (a : A) : Bialgebra.TensorProduct.rid R A (a ⊗ₜ r) = r • a := rfl + +@[simp] +theorem rid_symm_apply (a : A) : (Bialgebra.TensorProduct.rid R A).symm a = a ⊗ₜ 1 := rfl + +end Bialgebra.TensorProduct +namespace BialgHom + +variable {R A B C : Type u} [CommRing R] [Ring A] [Ring B] [Ring C] + [Bialgebra R A] [Bialgebra R B] [Bialgebra R C] + +variable (A) + +/-- `lTensor A f : A ⊗ B →ₐc A ⊗ C` is the natural bialgebra morphism induced by `f : B →ₐc C`. -/ +noncomputable abbrev lTensor (f : B →ₐc[R] C) : A ⊗[R] B →ₐc[R] A ⊗[R] C := + Bialgebra.TensorProduct.map (BialgHom.id R A) f + +/-- `rTensor A f : B ⊗ A →ₐc C ⊗ A` is the natural bialgebra morphism induced by `f : B →ₐc C`. -/ +noncomputable abbrev rTensor (f : B →ₐc[R] C) : B ⊗[R] A →ₐc[R] C ⊗[R] A := + Bialgebra.TensorProduct.map f (BialgHom.id R A) + +end BialgHom diff --git a/Mathlib/RingTheory/TensorProduct/Basic.lean b/Mathlib/RingTheory/TensorProduct/Basic.lean index dbb11acf62992..15539b8a077a5 100644 --- a/Mathlib/RingTheory/TensorProduct/Basic.lean +++ b/Mathlib/RingTheory/TensorProduct/Basic.lean @@ -956,6 +956,55 @@ theorem congr_trans [Algebra S C] [IsScalarTower R S C] theorem congr_symm (f : A ≃ₐ[S] B) (g : C ≃ₐ[R] D) : congr f.symm g.symm = (congr f g).symm := rfl +variable (R A B C) in +/-- Tensor product of algebras analogue of `mul_left_comm`. + +This is the algebra version of `TensorProduct.leftComm`. -/ +def leftComm : A ⊗[R] B ⊗[R] C ≃ₐ[R] B ⊗[R] A ⊗[R] C := + let e₁ := (Algebra.TensorProduct.assoc R A B C).symm + let e₂ := congr (Algebra.TensorProduct.comm R A B) (1 : C ≃ₐ[R] C) + let e₃ := Algebra.TensorProduct.assoc R B A C + e₁.trans (e₂.trans e₃) + +@[simp] +theorem leftComm_tmul (m : A) (n : B) (p : C) : + leftComm R A B C (m ⊗ₜ (n ⊗ₜ p)) = n ⊗ₜ (m ⊗ₜ p) := + rfl + +@[simp] +theorem leftComm_symm_tmul (m : A) (n : B) (p : C) : + (leftComm R A B C).symm (n ⊗ₜ (m ⊗ₜ p)) = m ⊗ₜ (n ⊗ₜ p) := + rfl + +@[simp] +theorem leftComm_toLinearEquiv : + (leftComm R A B C : _ ≃ₗ[R] _) = _root_.TensorProduct.leftComm R A B C := rfl + +variable (R A B C D) in +/-- Tensor product of algebras analogue of `mul_mul_mul_comm`. + +This is the algebra version of `TensorProduct.tensorTensorTensorComm`. -/ +def tensorTensorTensorComm : (A ⊗[R] B) ⊗[R] C ⊗[R] D ≃ₐ[R] (A ⊗[R] C) ⊗[R] B ⊗[R] D := + let e₁ := Algebra.TensorProduct.assoc R A B (C ⊗[R] D) + let e₂ := congr (1 : A ≃ₐ[R] A) (leftComm R B C D) + let e₃ := (Algebra.TensorProduct.assoc R A C (B ⊗[R] D)).symm + e₁.trans (e₂.trans e₃) + +@[simp] +theorem tensorTensorTensorComm_tmul (m : A) (n : B) (p : C) (q : D) : + tensorTensorTensorComm R A B C D (m ⊗ₜ n ⊗ₜ (p ⊗ₜ q)) = m ⊗ₜ p ⊗ₜ (n ⊗ₜ q) := + rfl + +@[simp] +theorem tensorTensorTensorComm_symm : + (tensorTensorTensorComm R A B C D).symm = tensorTensorTensorComm R A C B D := by + ext; rfl + +@[simp] +theorem tensorTensorTensorComm_toLinearEquiv : + (tensorTensorTensorComm R A B C D : _ ≃ₗ[R] _) = + _root_.TensorProduct.tensorTensorTensorComm R A B C D := rfl + end end Monoidal From 64e385171605476900ee7fb11ed510c7c83e6209 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 27 Oct 2024 19:58:55 +0000 Subject: [PATCH 485/505] feat: a product of indicators is the indicator of the product (#17964) --- Mathlib/Algebra/BigOperators/Pi.lean | 32 +++++++++++++++++++++++++- Mathlib/Algebra/BigOperators/Ring.lean | 6 +++++ Mathlib/Data/Finset/Lattice.lean | 20 ++++++++++------ 3 files changed, 50 insertions(+), 8 deletions(-) diff --git a/Mathlib/Algebra/BigOperators/Pi.lean b/Mathlib/Algebra/BigOperators/Pi.lean index 4f4ccb8bfbffe..215b6cb5305d4 100644 --- a/Mathlib/Algebra/BigOperators/Pi.lean +++ b/Mathlib/Algebra/BigOperators/Pi.lean @@ -3,7 +3,7 @@ Copyright (c) 2018 Simon Hudon. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Simon Hudon, Patrick Massot -/ -import Mathlib.Algebra.BigOperators.Group.Finset +import Mathlib.Algebra.BigOperators.GroupWithZero.Finset import Mathlib.Algebra.Group.Action.Pi import Mathlib.Algebra.Ring.Pi @@ -14,6 +14,9 @@ This file contains theorems relevant to big operators in binary and arbitrary pr of monoids and groups -/ +open scoped Finset + +variable {ι κ α : Type*} namespace Pi @@ -57,6 +60,33 @@ theorem pi_eq_sum_univ {ι : Type*} [Fintype ι] [DecidableEq ι] {R : Type*} [S ext simp +section CommSemiring +variable [CommSemiring α] + +lemma prod_indicator_apply (s : Finset ι) (f : ι → Set κ) (g : ι → κ → α) (j : κ) : + ∏ i ∈ s, (f i).indicator (g i) j = (⋂ x ∈ s, f x).indicator (∏ i ∈ s, g i) j := by + rw [Set.indicator] + split_ifs with hj + · rw [Finset.prod_apply] + congr! 1 with i hi + simp only [Finset.inf_set_eq_iInter, Set.mem_iInter] at hj + exact Set.indicator_of_mem (hj _ hi) _ + · obtain ⟨i, hi, hj⟩ := by simpa using hj + exact Finset.prod_eq_zero hi <| Set.indicator_of_not_mem hj _ + +lemma prod_indicator (s : Finset ι) (f : ι → Set κ) (g : ι → κ → α) : + ∏ i ∈ s, (f i).indicator (g i) = (⋂ x ∈ s, f x).indicator (∏ i ∈ s, g i) := by + ext a; simpa using prod_indicator_apply .. + +lemma prod_indicator_const_apply (s : Finset ι) (f : ι → Set κ) (g : κ → α) (j : κ) : + ∏ i ∈ s, (f i).indicator g j = (⋂ x ∈ s, f x).indicator (g ^ #s) j := by + simp [prod_indicator_apply] + +lemma prod_indicator_const (s : Finset ι) (f : ι → Set κ) (g : κ → α) : + ∏ i ∈ s, (f i).indicator g = (⋂ x ∈ s, f x).indicator (g ^ #s) := by simp [prod_indicator] + +end CommSemiring + section MulSingle variable {I : Type*} [DecidableEq I] {Z : I → Type*} diff --git a/Mathlib/Algebra/BigOperators/Ring.lean b/Mathlib/Algebra/BigOperators/Ring.lean index 650b3b139aa99..a25c14300bd4b 100644 --- a/Mathlib/Algebra/BigOperators/Ring.lean +++ b/Mathlib/Algebra/BigOperators/Ring.lean @@ -227,6 +227,12 @@ end CommSemiring section CommRing variable [CommRing α] +/-- The product of `f i - g i` over all of `s` is the sum over the powerset of `s` of the product of +`g` over a subset `t` times the product of `f` over the complement of `t` times `(-1) ^ #t`. -/ +lemma prod_sub [DecidableEq ι] (f g : ι → α) (s : Finset ι) : + ∏ i ∈ s, (f i - g i) = ∑ t ∈ s.powerset, (-1) ^ #t * (∏ i ∈ s \ t, f i) * ∏ i ∈ t, g i := by + simp [sub_eq_neg_add, prod_add, ← prod_const, ← prod_mul_distrib, mul_right_comm] + /-- `∏ i, (f i - g i) = (∏ i, f i) - ∑ i, g i * (∏ j < i, f j - g j) * (∏ j > i, f j)`. -/ lemma prod_sub_ordered [LinearOrder ι] (s : Finset ι) (f g : ι → α) : ∏ i ∈ s, (f i - g i) = diff --git a/Mathlib/Data/Finset/Lattice.lean b/Mathlib/Data/Finset/Lattice.lean index ad15a3cd7ffe6..154e0fd34f188 100644 --- a/Mathlib/Data/Finset/Lattice.lean +++ b/Mathlib/Data/Finset/Lattice.lean @@ -1181,12 +1181,18 @@ theorem mem_sup {α β} [DecidableEq β] {s : Finset α} {f : α → Multiset β end Multiset namespace Finset +variable [DecidableEq α] {s : Finset ι} {f : ι → Finset α} {a : α} -theorem mem_sup {α β} [DecidableEq β] {s : Finset α} {f : α → Finset β} {x : β} : - x ∈ s.sup f ↔ ∃ v ∈ s, x ∈ f v := by - change _ ↔ ∃ v ∈ s, x ∈ (f v).val - rw [← Multiset.mem_sup, ← Multiset.mem_toFinset, sup_toFinset] - simp_rw [val_toFinset] +set_option linter.docPrime false in +@[simp] lemma mem_sup' (hs) : a ∈ s.sup' hs f ↔ ∃ i ∈ s, a ∈ f i := by + induction' hs using Nonempty.cons_induction <;> simp [*] + +set_option linter.docPrime false in +@[simp] lemma mem_inf' (hs) : a ∈ s.inf' hs f ↔ ∀ i ∈ s, a ∈ f i := by + induction' hs using Nonempty.cons_induction <;> simp [*] + +@[simp] lemma mem_sup : a ∈ s.sup f ↔ ∃ i ∈ s, a ∈ f i := by + induction' s using cons_induction <;> simp [*] theorem sup_eq_biUnion {α β} [DecidableEq β] (s : Finset α) (t : α → Finset β) : s.sup t = s.biUnion t := by @@ -1194,14 +1200,14 @@ theorem sup_eq_biUnion {α β} [DecidableEq β] (s : Finset α) (t : α → Fins rw [mem_sup, mem_biUnion] @[simp] -theorem sup_singleton'' [DecidableEq α] (s : Finset β) (f : β → α) : +theorem sup_singleton'' (s : Finset β) (f : β → α) : (s.sup fun b => {f b}) = s.image f := by ext a rw [mem_sup, mem_image] simp only [mem_singleton, eq_comm] @[simp] -theorem sup_singleton' [DecidableEq α] (s : Finset α) : s.sup singleton = s := +theorem sup_singleton' (s : Finset α) : s.sup singleton = s := (s.sup_singleton'' _).trans image_id end Finset From 1dd053cf4b4ced5c422f6471e227eb19c261394a Mon Sep 17 00:00:00 2001 From: mathlib4-update-dependencies-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Sun, 27 Oct 2024 19:58:56 +0000 Subject: [PATCH 486/505] chore: update Mathlib dependencies 2024-10-27 (#18297) This PR updates the Mathlib dependencies. --- lake-manifest.json | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lake-manifest.json b/lake-manifest.json index 7a31bfbacb159..cb37bf904e85f 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,7 +5,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "dc72dcdb8e97b3c56bd70f06f043ed2dee3258e6", + "rev": "4d2cb85d87e0e728520fe0bf531cfdddb78fde7a", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", From feac7b36ef96a86aa423e04b63146f905a6280f3 Mon Sep 17 00:00:00 2001 From: Christopher Hoskin Date: Sun, 27 Oct 2024 20:35:08 +0000 Subject: [PATCH 487/505] feature(Analysis/LocallyConvex/AbsConvex): The absolutely convex hull of a totally bounded set is totally bounded (#17204) In a locally convex space, the absolutely convex hull of a totally bounded set is totally bounded. The proof given here follows Proposition 3 of Bourbaki TVS II.25. Co-authored-by: Christopher Hoskin --- Mathlib.lean | 1 + Mathlib/Analysis/Convex/Hull.lean | 5 ++ Mathlib/Analysis/Convex/TotallyBounded.lean | 52 +++++++++++++++++++ Mathlib/Analysis/LocallyConvex/AbsConvex.lean | 13 +++++ Mathlib/Topology/Algebra/UniformGroup.lean | 5 ++ 5 files changed, 76 insertions(+) create mode 100644 Mathlib/Analysis/Convex/TotallyBounded.lean diff --git a/Mathlib.lean b/Mathlib.lean index 8613416a68ff8..99ff2c3fb06df 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1197,6 +1197,7 @@ import Mathlib.Analysis.Convex.StrictConvexBetween import Mathlib.Analysis.Convex.StrictConvexSpace import Mathlib.Analysis.Convex.Strong import Mathlib.Analysis.Convex.Topology +import Mathlib.Analysis.Convex.TotallyBounded import Mathlib.Analysis.Convex.Uniform import Mathlib.Analysis.Convolution import Mathlib.Analysis.Distribution.AEEqOfIntegralContDiff diff --git a/Mathlib/Analysis/Convex/Hull.lean b/Mathlib/Analysis/Convex/Hull.lean index 996ebbb415ab1..846de50849e76 100644 --- a/Mathlib/Analysis/Convex/Hull.lean +++ b/Mathlib/Analysis/Convex/Hull.lean @@ -148,6 +148,11 @@ theorem LinearMap.image_convexHull (f : E →ₗ[𝕜] F) (s : Set E) : f '' convexHull 𝕜 s = convexHull 𝕜 (f '' s) := f.isLinear.image_convexHull s +theorem convexHull_add_subset {s t : Set E} : + convexHull 𝕜 (s + t) ⊆ convexHull 𝕜 s + convexHull 𝕜 t := + convexHull_min (add_subset_add (subset_convexHull _ _) (subset_convexHull _ _)) + (Convex.add (convex_convexHull 𝕜 s) (convex_convexHull 𝕜 t)) + end AddCommMonoid end OrderedSemiring diff --git a/Mathlib/Analysis/Convex/TotallyBounded.lean b/Mathlib/Analysis/Convex/TotallyBounded.lean new file mode 100644 index 0000000000000..5506f8e7dced5 --- /dev/null +++ b/Mathlib/Analysis/Convex/TotallyBounded.lean @@ -0,0 +1,52 @@ +/- +Copyright (c) 2024 Christopher Hoskin. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Christopher Hoskin +-/ + +import Mathlib.Topology.UniformSpace.Cauchy +import Mathlib.Analysis.Convex.Hull +import Mathlib.Topology.Algebra.UniformGroup +import Mathlib.Topology.Algebra.Module.LocallyConvex + +/-! +# Totally Bounded sets and Convex Hulls + +## Main statements + +- `totallyBounded_convexHull` The convex hull of a totally bounded set is totally bounded. + +## References + +* [Bourbaki, *Topological Vector Spaces*][bourbaki1987] + +## Tags + +convex, totally bounded +-/ + +open Set Pointwise + +variable (E : Type*) {s : Set E} +variable [AddCommGroup E] [Module ℝ E] +variable [UniformSpace E] [UniformAddGroup E] [lcs : LocallyConvexSpace ℝ E] [ContinuousSMul ℝ E] + +theorem totallyBounded_convexHull (hs : TotallyBounded s) : + TotallyBounded (convexHull ℝ s) := by + rw [totallyBounded_iff_subset_finite_iUnion_nhds_zero] + intro U hU + obtain ⟨W, hW₁, hW₂⟩ := exists_nhds_zero_half hU + obtain ⟨V, ⟨hV₁, hV₂, hV₃⟩⟩ := (locallyConvexSpace_iff_exists_convex_subset_zero ℝ E).mp lcs W hW₁ + obtain ⟨t, ⟨htf, hts⟩⟩ := (totallyBounded_iff_subset_finite_iUnion_nhds_zero.mp hs) _ hV₁ + obtain ⟨t', ⟨htf', hts'⟩⟩ := (totallyBounded_iff_subset_finite_iUnion_nhds_zero.mp + (IsCompact.totallyBounded (Finite.isCompact_convexHull htf)) _ hV₁) + use t', htf' + simp only [iUnion_vadd_set, vadd_eq_add] at hts hts' ⊢ + calc convexHull ℝ s + _ ⊆ convexHull ℝ (t + V) := convexHull_mono hts + _ ⊆ convexHull ℝ t + convexHull ℝ V := convexHull_add_subset + _ = convexHull ℝ t + V := by rw [hV₂.convexHull_eq] + _ ⊆ t' + V + V := add_subset_add_right hts' + _ = t' + (V + V) := by rw [add_assoc] + _ ⊆ t' + (W + W) := add_subset_add_left (add_subset_add hV₃ hV₃) + _ ⊆ t' + U := add_subset_add_left (add_subset_iff.mpr hW₂) diff --git a/Mathlib/Analysis/LocallyConvex/AbsConvex.lean b/Mathlib/Analysis/LocallyConvex/AbsConvex.lean index f9e44391da1b4..c13ab3c9c7b85 100644 --- a/Mathlib/Analysis/LocallyConvex/AbsConvex.lean +++ b/Mathlib/Analysis/LocallyConvex/AbsConvex.lean @@ -6,6 +6,7 @@ Authors: Moritz Doll import Mathlib.Analysis.LocallyConvex.BalancedCoreHull import Mathlib.Analysis.LocallyConvex.WithSeminorms import Mathlib.Analysis.Convex.Gauge +import Mathlib.Analysis.Convex.TotallyBounded /-! # Absolutely convex sets @@ -238,6 +239,18 @@ theorem convexHull_union_neg_eq_absConvexHull {s : Set E} : rw [← Convex.convexHull_eq (convex_convexHull ℝ (s ∪ -s))] exact convexHull_mono balancedHull_subset_convexHull_union_neg) +variable (E 𝕜) {s : Set E} +variable [NontriviallyNormedField 𝕜] [Module 𝕜 E] [SMulCommClass ℝ 𝕜 E] +variable [UniformSpace E] [UniformAddGroup E] [lcs : LocallyConvexSpace ℝ E] [ContinuousSMul ℝ E] + +-- TVS II.25 Prop3 +theorem totallyBounded_absConvexHull (hs : TotallyBounded s) : + TotallyBounded (absConvexHull ℝ s) := by + rw [← convexHull_union_neg_eq_absConvexHull] + apply totallyBounded_convexHull + rw [totallyBounded_union] + exact ⟨hs, totallyBounded_neg hs⟩ + end section AbsolutelyConvexSets diff --git a/Mathlib/Topology/Algebra/UniformGroup.lean b/Mathlib/Topology/Algebra/UniformGroup.lean index 1c3bc90dfda28..379ce246257ff 100644 --- a/Mathlib/Topology/Algebra/UniformGroup.lean +++ b/Mathlib/Topology/Algebra/UniformGroup.lean @@ -417,6 +417,11 @@ theorem totallyBounded_iff_subset_finite_iUnion_nhds_one {s : Set α} : (𝓝 (1 : α)).basis_sets.uniformity_of_nhds_one_inv_mul_swapped.totallyBounded_iff.trans <| by simp [← preimage_smul_inv, preimage] +@[to_additive] +theorem totallyBounded_inv {s : Set α} (hs : TotallyBounded s) : TotallyBounded (s⁻¹) := by + convert TotallyBounded.image hs uniformContinuous_inv + aesop + section UniformConvergence variable {ι : Type*} {l : Filter ι} {l' : Filter β} {f f' : ι → β → α} {g g' : β → α} {s : Set β} From c57a78f1e6e36f9af9e045f226de6c084db40e63 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sun, 27 Oct 2024 20:54:24 +0000 Subject: [PATCH 488/505] feat(Topology/Algebra): generalize `CompleteSpace` instances (#17244) --- .../Algebra/Module/Alternating/Topology.lean | 17 +++++- Mathlib/Topology/Algebra/Module/Basic.lean | 10 ++++ .../Algebra/Module/Multilinear/Topology.lean | 34 ++++++++--- .../Algebra/Module/StrongTopology.lean | 59 +++++++++++++++---- 4 files changed, 102 insertions(+), 18 deletions(-) diff --git a/Mathlib/Topology/Algebra/Module/Alternating/Topology.lean b/Mathlib/Topology/Algebra/Module/Alternating/Topology.lean index 61776f9e7c9c2..6002b8905bc13 100644 --- a/Mathlib/Topology/Algebra/Module/Alternating/Topology.lean +++ b/Mathlib/Topology/Algebra/Module/Alternating/Topology.lean @@ -77,13 +77,28 @@ instance instUniformContinuousConstSMul {M : Type*} UniformContinuousConstSMul M (E [⋀^ι]→L[𝕜] F) := isUniformEmbedding_toContinuousMultilinearMap.uniformContinuousConstSMul fun _ _ ↦ rfl +theorem isUniformInducing_postcomp {G : Type*} [AddCommGroup G] [UniformSpace G] [UniformAddGroup G] + [Module 𝕜 G] (g : F →L[𝕜] G) (hg : IsUniformInducing g) : + IsUniformInducing (g.compContinuousAlternatingMap : (E [⋀^ι]→L[𝕜] F) → (E [⋀^ι]→L[𝕜] G)) := by + rw [← isUniformEmbedding_toContinuousMultilinearMap.1.of_comp_iff] + exact (ContinuousMultilinearMap.isUniformInducing_postcomp g hg).comp + isUniformEmbedding_toContinuousMultilinearMap.1 + section CompleteSpace -variable [ContinuousSMul 𝕜 E] [ContinuousConstSMul 𝕜 F] [CompleteSpace F] [T2Space F] +variable [ContinuousSMul 𝕜 E] [ContinuousConstSMul 𝕜 F] [CompleteSpace F] open UniformOnFun in theorem completeSpace (h : RestrictGenTopology {s : Set (ι → E) | IsVonNBounded 𝕜 s}) : CompleteSpace (E [⋀^ι]→L[𝕜] F) := by + wlog hF : T2Space F generalizing F + · rw [(isUniformInducing_postcomp (SeparationQuotient.mkCLM _ _) + SeparationQuotient.isUniformInducing_mk).completeSpace_congr] + · exact this inferInstance + · intro f + use (SeparationQuotient.outCLM _ _).compContinuousAlternatingMap f + ext + simp have := ContinuousMultilinearMap.completeSpace (F := F) h rw [completeSpace_iff_isComplete_range isUniformEmbedding_toContinuousMultilinearMap.isUniformInducing] diff --git a/Mathlib/Topology/Algebra/Module/Basic.lean b/Mathlib/Topology/Algebra/Module/Basic.lean index fbc19334443c1..70d2d820610b9 100644 --- a/Mathlib/Topology/Algebra/Module/Basic.lean +++ b/Mathlib/Topology/Algebra/Module/Basic.lean @@ -429,6 +429,16 @@ theorem coe_copy (f : M₁ →SL[σ₁₂] M₂) (f' : M₁ → M₂) (h : f' = theorem copy_eq (f : M₁ →SL[σ₁₂] M₂) (f' : M₁ → M₂) (h : f' = ⇑f) : f.copy f' h = f := DFunLike.ext' h +theorem range_coeFn_eq : + Set.range ((⇑) : (M₁ →SL[σ₁₂] M₂) → (M₁ → M₂)) = + {f | Continuous f} ∩ Set.range ((⇑) : (M₁ →ₛₗ[σ₁₂] M₂) → (M₁ → M₂)) := by + ext f + constructor + · rintro ⟨f, rfl⟩ + exact ⟨f.continuous, f, rfl⟩ + · rintro ⟨hfc, f, rfl⟩ + exact ⟨⟨f, hfc⟩, rfl⟩ + -- make some straightforward lemmas available to `simp`. protected theorem map_zero (f : M₁ →SL[σ₁₂] M₂) : f (0 : M₁) = 0 := map_zero f diff --git a/Mathlib/Topology/Algebra/Module/Multilinear/Topology.lean b/Mathlib/Topology/Algebra/Module/Multilinear/Topology.lean index 19af25db016c5..87d86ced20d2e 100644 --- a/Mathlib/Topology/Algebra/Module/Multilinear/Topology.lean +++ b/Mathlib/Topology/Algebra/Module/Multilinear/Topology.lean @@ -5,6 +5,7 @@ Authors: Yury Kudryashov -/ import Mathlib.Topology.Algebra.Module.Multilinear.Bounded import Mathlib.Topology.Algebra.Module.UniformConvergence +import Mathlib.Topology.Algebra.SeparationQuotient.Section import Mathlib.Topology.Hom.ContinuousEvalConst /-! @@ -66,15 +67,20 @@ section UniformAddGroup variable [UniformSpace F] [UniformAddGroup F] +lemma isUniformInducing_toUniformOnFun : + IsUniformInducing (toUniformOnFun : + ContinuousMultilinearMap 𝕜 E F → ((Π i, E i) →ᵤ[{s | IsVonNBounded 𝕜 s}] F)) := ⟨rfl⟩ + lemma isUniformEmbedding_toUniformOnFun : - IsUniformEmbedding (toUniformOnFun : ContinuousMultilinearMap 𝕜 E F → _) where - inj := DFunLike.coe_injective - comap_uniformity := rfl + IsUniformEmbedding (toUniformOnFun : ContinuousMultilinearMap 𝕜 E F → _) := + ⟨isUniformInducing_toUniformOnFun, DFunLike.coe_injective⟩ @[deprecated (since := "2024-10-01")] alias uniformEmbedding_toUniformOnFun := isUniformEmbedding_toUniformOnFun -lemma embedding_toUniformOnFun : Embedding (toUniformOnFun : ContinuousMultilinearMap 𝕜 E F → _) := +lemma embedding_toUniformOnFun : + Embedding (toUniformOnFun : ContinuousMultilinearMap 𝕜 E F → + ((Π i, E i) →ᵤ[{s | IsVonNBounded 𝕜 s}] F)) := isUniformEmbedding_toUniformOnFun.embedding theorem uniformContinuous_coe_fun [∀ i, ContinuousSMul 𝕜 (E i)] : @@ -97,19 +103,33 @@ instance instUniformContinuousConstSMul {M : Type*} haveI := uniformContinuousConstSMul_of_continuousConstSMul M F isUniformEmbedding_toUniformOnFun.uniformContinuousConstSMul fun _ _ ↦ rfl +theorem isUniformInducing_postcomp + {G : Type*} [AddCommGroup G] [UniformSpace G] [UniformAddGroup G] [Module 𝕜 G] + (g : F →L[𝕜] G) (hg : IsUniformInducing g) : + IsUniformInducing (g.compContinuousMultilinearMap : + ContinuousMultilinearMap 𝕜 E F → ContinuousMultilinearMap 𝕜 E G) := by + rw [← isUniformInducing_toUniformOnFun.of_comp_iff] + exact (UniformOnFun.postcomp_isUniformInducing hg).comp isUniformInducing_toUniformOnFun + section CompleteSpace -variable [∀ i, ContinuousSMul 𝕜 (E i)] [ContinuousConstSMul 𝕜 F] [CompleteSpace F] [T2Space F] +variable [∀ i, ContinuousSMul 𝕜 (E i)] [ContinuousConstSMul 𝕜 F] [CompleteSpace F] open UniformOnFun in theorem completeSpace (h : RestrictGenTopology {s : Set (Π i, E i) | IsVonNBounded 𝕜 s}) : CompleteSpace (ContinuousMultilinearMap 𝕜 E F) := by classical + wlog hF : T2Space F generalizing F + · rw [(isUniformInducing_postcomp (SeparationQuotient.mkCLM _ _) + SeparationQuotient.isUniformInducing_mk).completeSpace_congr] + · exact this inferInstance + · intro f + use (SeparationQuotient.outCLM _ _).compContinuousMultilinearMap f + simp [DFunLike.ext_iff] have H : ∀ {m : Π i, E i}, Continuous fun f : (Π i, E i) →ᵤ[{s | IsVonNBounded 𝕜 s}] F ↦ toFun _ f m := (uniformContinuous_eval (isVonNBounded_covers) _).continuous - rw [completeSpace_iff_isComplete_range isUniformEmbedding_toUniformOnFun.isUniformInducing, - range_toUniformOnFun] + rw [completeSpace_iff_isComplete_range isUniformInducing_toUniformOnFun, range_toUniformOnFun] simp only [setOf_and, setOf_forall] apply_rules [IsClosed.isComplete, IsClosed.inter] · exact UniformOnFun.isClosed_setOf_continuous h diff --git a/Mathlib/Topology/Algebra/Module/StrongTopology.lean b/Mathlib/Topology/Algebra/Module/StrongTopology.lean index 3132c5f3d16b8..40a009bdf15e6 100644 --- a/Mathlib/Topology/Algebra/Module/StrongTopology.lean +++ b/Mathlib/Topology/Algebra/Module/StrongTopology.lean @@ -1,9 +1,10 @@ /- Copyright (c) 2022 Anatole Dedecker. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Anatole Dedecker +Authors: Anatole Dedecker, Yury Kudryashov -/ import Mathlib.Topology.Algebra.Module.UniformConvergence +import Mathlib.Topology.Algebra.SeparationQuotient.Section import Mathlib.Topology.Hom.ContinuousEvalConst /-! @@ -55,7 +56,6 @@ sets). uniform convergence, bounded convergence -/ - open scoped Topology UniformConvergence Uniformity open Filter Set Function Bornology @@ -75,17 +75,16 @@ uniform convergence on the elements of `𝔖`". If the continuous linear image of any element of `𝔖` is bounded, this makes `E →SL[σ] F` a topological vector space. -/ @[nolint unusedArguments] -def UniformConvergenceCLM [TopologicalSpace F] [TopologicalAddGroup F] (_ : Set (Set E)) := - E →SL[σ] F +def UniformConvergenceCLM [TopologicalSpace F] (_ : Set (Set E)) := E →SL[σ] F namespace UniformConvergenceCLM -instance instFunLike [TopologicalSpace F] [TopologicalAddGroup F] - (𝔖 : Set (Set E)) : FunLike (UniformConvergenceCLM σ F 𝔖) E F := +instance instFunLike [TopologicalSpace F] (𝔖 : Set (Set E)) : + FunLike (UniformConvergenceCLM σ F 𝔖) E F := ContinuousLinearMap.funLike -instance instContinuousSemilinearMapClass [TopologicalSpace F] [TopologicalAddGroup F] - (𝔖 : Set (Set E)) : ContinuousSemilinearMapClass (UniformConvergenceCLM σ F 𝔖) σ E F := +instance instContinuousSemilinearMapClass [TopologicalSpace F] (𝔖 : Set (Set E)) : + ContinuousSemilinearMapClass (UniformConvergenceCLM σ F 𝔖) σ E F := ContinuousLinearMap.continuousSemilinearMapClass instance instTopologicalSpace [TopologicalSpace F] [TopologicalAddGroup F] (𝔖 : Set (Set E)) : @@ -120,9 +119,13 @@ theorem uniformity_toTopologicalSpace_eq [UniformSpace F] [UniformAddGroup F] ( UniformConvergenceCLM.instTopologicalSpace σ F 𝔖 := rfl +theorem isUniformInducing_coeFn [UniformSpace F] [UniformAddGroup F] (𝔖 : Set (Set E)) : + IsUniformInducing (α := UniformConvergenceCLM σ F 𝔖) (UniformOnFun.ofFun 𝔖 ∘ DFunLike.coe) := + ⟨rfl⟩ + theorem isUniformEmbedding_coeFn [UniformSpace F] [UniformAddGroup F] (𝔖 : Set (Set E)) : IsUniformEmbedding (α := UniformConvergenceCLM σ F 𝔖) (UniformOnFun.ofFun 𝔖 ∘ DFunLike.coe) := - ⟨⟨rfl⟩, DFunLike.coe_injective⟩ + ⟨isUniformInducing_coeFn .., DFunLike.coe_injective⟩ @[deprecated (since := "2024-10-01")] alias uniformEmbedding_coeFn := isUniformEmbedding_coeFn @@ -262,7 +265,7 @@ instance instUniformContinuousConstSMul (M : Type*) [Monoid M] [DistribMulAction M F] [SMulCommClass 𝕜₂ M F] [UniformSpace F] [UniformAddGroup F] [UniformContinuousConstSMul M F] (𝔖 : Set (Set E)) : UniformContinuousConstSMul M (UniformConvergenceCLM σ F 𝔖) := - (isUniformEmbedding_coeFn σ F 𝔖).isUniformInducing.uniformContinuousConstSMul fun _ _ ↦ by rfl + (isUniformInducing_coeFn σ F 𝔖).uniformContinuousConstSMul fun _ _ ↦ by rfl instance instContinuousConstSMul (M : Type*) [Monoid M] [DistribMulAction M F] [SMulCommClass 𝕜₂ M F] @@ -280,6 +283,32 @@ theorem tendsto_iff_tendstoUniformlyOn {ι : Type*} {p : Filter ι} [UniformSpac rw [(embedding_coeFn σ F 𝔖).tendsto_nhds_iff, UniformOnFun.tendsto_iff_tendstoUniformlyOn] rfl +variable {F} in +theorem isUniformInducing_postcomp + {G : Type*} [AddCommGroup G] [UniformSpace G] [UniformAddGroup G] + {𝕜₃ : Type*} [NormedField 𝕜₃] [Module 𝕜₃ G] + {τ : 𝕜₂ →+* 𝕜₃} {ρ : 𝕜₁ →+* 𝕜₃} [RingHomCompTriple σ τ ρ] [UniformSpace F] [UniformAddGroup F] + (g : F →SL[τ] G) (hg : IsUniformInducing g) (𝔖 : Set (Set E)) : + IsUniformInducing (α := UniformConvergenceCLM σ F 𝔖) (β := UniformConvergenceCLM ρ G 𝔖) + g.comp := by + rw [← (isUniformInducing_coeFn _ _ _).of_comp_iff] + exact (UniformOnFun.postcomp_isUniformInducing hg).comp (isUniformInducing_coeFn _ _ _) + +theorem completeSpace [UniformSpace F] [UniformAddGroup F] [ContinuousSMul 𝕜₂ F] [CompleteSpace F] + {𝔖 : Set (Set E)} (h𝔖 : RestrictGenTopology 𝔖) (h𝔖U : ⋃₀ 𝔖 = univ) : + CompleteSpace (UniformConvergenceCLM σ F 𝔖) := by + wlog hF : T2Space F generalizing F + · rw [(isUniformInducing_postcomp σ (SeparationQuotient.mkCLM 𝕜₂ F) + SeparationQuotient.isUniformInducing_mk _).completeSpace_congr] + exacts [this _ inferInstance, SeparationQuotient.postcomp_mkCLM_surjective F σ E] + rw [completeSpace_iff_isComplete_range (isUniformInducing_coeFn _ _ _)] + apply IsClosed.isComplete + have H₁ : IsClosed {f : E →ᵤ[𝔖] F | Continuous ((UniformOnFun.toFun 𝔖) f)} := + UniformOnFun.isClosed_setOf_continuous h𝔖 + convert H₁.inter <| (LinearMap.isClosed_range_coe E F σ).preimage + (UniformOnFun.uniformContinuous_toFun h𝔖U).continuous + exact ContinuousLinearMap.range_coeFn_eq + variable {𝔖₁ 𝔖₂ : Set (Set E)} theorem uniformSpace_mono [UniformSpace F] [UniformAddGroup F] (h : 𝔖₂ ⊆ 𝔖₁) : @@ -415,6 +444,16 @@ theorem isVonNBounded_iff {R : Type*} [NormedDivisionRing R] ∀ s, IsVonNBounded 𝕜₁ s → IsVonNBounded R (Set.image2 (fun f x ↦ f x) S s) := UniformConvergenceCLM.isVonNBounded_iff +theorem completeSpace [UniformSpace F] [UniformAddGroup F] [ContinuousSMul 𝕜₂ F] [CompleteSpace F] + [ContinuousSMul 𝕜₁ E] (h : RestrictGenTopology {s : Set E | IsVonNBounded 𝕜₁ s}) : + CompleteSpace (E →SL[σ] F) := + UniformConvergenceCLM.completeSpace _ _ h isVonNBounded_covers + +instance instCompleteSpace [TopologicalAddGroup E] [ContinuousSMul 𝕜₁ E] [SequentialSpace E] + [UniformSpace F] [UniformAddGroup F] [ContinuousSMul 𝕜₂ F] [CompleteSpace F] : + CompleteSpace (E →SL[σ] F) := + completeSpace <| .of_seq fun _ _ h ↦ (h.isVonNBounded_range 𝕜₁).insert _ + variable (G) [TopologicalSpace F] [TopologicalSpace G] /-- Pre-composition by a *fixed* continuous linear map as a continuous linear map. From 380f469e41f87c305b0706bc38dc1b785603e481 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Sun, 27 Oct 2024 21:51:06 +0000 Subject: [PATCH 489/505] feat: isometric continuous functional calculi (#17837) --- Mathlib.lean | 1 + .../Instances.lean | 20 + .../Isometric.lean | 488 ++++++++++++++++++ Mathlib/Analysis/Normed/Algebra/Spectrum.lean | 17 + .../Analysis/Normed/Algebra/Unitization.lean | 4 + 5 files changed, 530 insertions(+) create mode 100644 Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Isometric.lean diff --git a/Mathlib.lean b/Mathlib.lean index 99ff2c3fb06df..cc6e5cead888e 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1005,6 +1005,7 @@ import Mathlib.Analysis.CStarAlgebra.Classes import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Basic import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Instances import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Integral +import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Isometric import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.NonUnital import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Order import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.PosPart diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean index b4cc643968f53..a2cba0b1582a9 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean @@ -169,10 +169,30 @@ instance IsStarNormal.instContinuousFunctionalCalculus {A : Type*} [CStarAlgebra AlgEquiv.spectrum_eq (continuousFunctionalCalculus a), ContinuousMap.spectrum_eq_range] case predicate_hom => exact fun f ↦ ⟨by rw [← map_star]; exact Commute.all (star f) f |>.map _⟩ +lemma cfcHom_eq_of_isStarNormal {A : Type*} [CStarAlgebra A] (a : A) [ha : IsStarNormal a] : + cfcHom ha = (elementalStarAlgebra ℂ a).subtype.comp (continuousFunctionalCalculus a) := by + refine cfcHom_eq_of_continuous_of_map_id ha _ ?_ ?_ + · exact continuous_subtype_val.comp <| + (StarAlgEquiv.isometry (continuousFunctionalCalculus a)).continuous + · simp [continuousFunctionalCalculus_map_id a] + instance IsStarNormal.instNonUnitalContinuousFunctionalCalculus {A : Type*} [NonUnitalCStarAlgebra A] : NonUnitalContinuousFunctionalCalculus ℂ (IsStarNormal : A → Prop) := RCLike.nonUnitalContinuousFunctionalCalculus Unitization.isStarNormal_inr +open Unitization in +lemma inr_comp_cfcₙHom_eq_cfcₙAux {A : Type*} [NonUnitalCStarAlgebra A] (a : A) + [ha : IsStarNormal a] : (inrNonUnitalStarAlgHom ℂ A).comp (cfcₙHom ha) = + cfcₙAux (isStarNormal_inr (R := ℂ) (A := A)) a ha := by + have h (a : A) := isStarNormal_inr (R := ℂ) (A := A) (a := a) + refine @UniqueNonUnitalContinuousFunctionalCalculus.eq_of_continuous_of_map_id + _ _ _ _ _ _ _ _ _ _ _ inferInstance inferInstance _ (σₙ ℂ a) _ _ rfl _ _ ?_ ?_ ?_ + · show Continuous (fun f ↦ (cfcₙHom ha f : Unitization ℂ A)); fun_prop + · exact isClosedEmbedding_cfcₙAux @(h) a ha |>.continuous + · trans (a : Unitization ℂ A) + · congrm(inr $(cfcₙHom_id ha)) + · exact cfcₙAux_id @(h) a ha |>.symm + end Normal /-! diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Isometric.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Isometric.lean new file mode 100644 index 0000000000000..0ab77c3d67964 --- /dev/null +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Isometric.lean @@ -0,0 +1,488 @@ +/- +Copyright (c) 2024 Jireh Loreaux. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jireh Loreaux +-/ +import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Instances + +/-! # Isometric continuous functional calculus + +This file adds a class for an *isometric* continuous functional calculus. This is separate from the +usual `ContinuousFunctionalCalculus` class because we prefer not to require a metric (or a norm) on +the algebra for reasons discussed in the module documentation for that file. + +Of course, with a metric on the algebra and an isometric continuous functional calculus, the +algebra must *be* a C⋆-algebra already. As such, it may seem like this class is not useful. However, +the main purpose is to allow for the continuous functional calculus to be a isometric for the other +scalar rings `ℝ` and `ℝ≥0` too. +-/ + +local notation "σ" => spectrum +local notation "σₙ" => quasispectrum + +/-! ### Isometric continuous functional calculus for unital algebras -/ +section Unital + +/-- An extension of the `ContinuousFunctionalCalculus` requiring that `cfcHom` is an isometry. -/ +class IsometricContinuousFunctionalCalculus (R A : Type*) (p : outParam (A → Prop)) + [CommSemiring R] [StarRing R] [MetricSpace R] [TopologicalSemiring R] [ContinuousStar R] + [Ring A] [StarRing A] [MetricSpace A] [Algebra R A] + extends ContinuousFunctionalCalculus R p : Prop where + isometric (a : A) (ha : p a) : Isometry (cfcHom ha (R := R)) + +section MetricSpace + +open scoped ContinuousFunctionalCalculus + +lemma isometry_cfcHom {R A : Type*} {p : outParam (A → Prop)} [CommSemiring R] [StarRing R] + [MetricSpace R] [TopologicalSemiring R] [ContinuousStar R] [Ring A] [StarRing A] + [MetricSpace A] [Algebra R A] [IsometricContinuousFunctionalCalculus R A p] + (a : A) (ha : p a := by cfc_tac) : + Isometry (cfcHom (show p a from ha) (R := R)) := + IsometricContinuousFunctionalCalculus.isometric a ha + +end MetricSpace + +section NormedRing + +open scoped ContinuousFunctionalCalculus + +variable {𝕜 A : Type*} {p : outParam (A → Prop)} +variable [RCLike 𝕜] [NormedRing A] [StarRing A] [NormedAlgebra 𝕜 A] +variable [IsometricContinuousFunctionalCalculus 𝕜 A p] + +lemma norm_cfcHom (a : A) (f : C(σ 𝕜 a, 𝕜)) (ha : p a := by cfc_tac) : + ‖cfcHom (show p a from ha) f‖ = ‖f‖ := by + refine isometry_cfcHom a |>.norm_map_of_map_zero (map_zero _) f + +lemma nnnorm_cfcHom (a : A) (f : C(σ 𝕜 a, 𝕜)) (ha : p a := by cfc_tac) : + ‖cfcHom (show p a from ha) f‖₊ = ‖f‖₊ := + Subtype.ext <| norm_cfcHom a f ha + +lemma IsGreatest.norm_cfc [Nontrivial A] (f : 𝕜 → 𝕜) (a : A) + (hf : ContinuousOn f (σ 𝕜 a) := by cfc_cont_tac) (ha : p a := by cfc_tac) : + IsGreatest ((fun x ↦ ‖f x‖) '' spectrum 𝕜 a) ‖cfc f a‖ := by + obtain ⟨x, hx⟩ := ContinuousFunctionalCalculus.isCompact_spectrum a + |>.image_of_continuousOn hf.norm |>.exists_isGreatest <| + (ContinuousFunctionalCalculus.spectrum_nonempty a ha).image _ + obtain ⟨x, hx', rfl⟩ := hx.1 + convert hx + rw [cfc_apply f a, norm_cfcHom a _] + apply le_antisymm + · apply ContinuousMap.norm_le _ (norm_nonneg _) |>.mpr + rintro ⟨y, hy⟩ + exact hx.2 ⟨y, hy, rfl⟩ + · exact le_trans (by simp) <| ContinuousMap.norm_coe_le_norm _ (⟨x, hx'⟩ : σ 𝕜 a) + +lemma IsGreatest.nnnorm_cfc [Nontrivial A] (f : 𝕜 → 𝕜) (a : A) + (hf : ContinuousOn f (σ 𝕜 a) := by cfc_cont_tac) (ha : p a := by cfc_tac) : + IsGreatest ((fun x ↦ ‖f x‖₊) '' σ 𝕜 a) ‖cfc f a‖₊ := by + convert Real.toNNReal_mono.map_isGreatest (.norm_cfc f a) + all_goals simp [Set.image_image, norm_toNNReal] + +lemma norm_apply_le_norm_cfc (f : 𝕜 → 𝕜) (a : A) ⦃x : 𝕜⦄ (hx : x ∈ σ 𝕜 a) + (hf : ContinuousOn f (σ 𝕜 a) := by cfc_cont_tac) (ha : p a := by cfc_tac) : + ‖f x‖ ≤ ‖cfc f a‖ := by + revert hx + nontriviality A + exact (IsGreatest.norm_cfc f a hf ha |>.2 ⟨x, ·, rfl⟩) + +lemma nnnorm_apply_le_nnnorm_cfc (f : 𝕜 → 𝕜) (a : A) ⦃x : 𝕜⦄ (hx : x ∈ σ 𝕜 a) + (hf : ContinuousOn f (σ 𝕜 a) := by cfc_cont_tac) (ha : p a := by cfc_tac) : + ‖f x‖₊ ≤ ‖cfc f a‖₊ := + norm_apply_le_norm_cfc f a hx + +lemma norm_cfc_le {f : 𝕜 → 𝕜} {a : A} {c : ℝ} (hc : 0 ≤ c) (h : ∀ x ∈ σ 𝕜 a, ‖f x‖ ≤ c) : + ‖cfc f a‖ ≤ c := by + obtain (_ | _) := subsingleton_or_nontrivial A + · simpa [Subsingleton.elim (cfc f a) 0] + · refine cfc_cases (‖·‖ ≤ c) a f (by simpa) fun hf ha ↦ ?_ + simp only [← cfc_apply f a, isLUB_le_iff (IsGreatest.norm_cfc f a hf ha |>.isLUB)] + rintro - ⟨x, hx, rfl⟩ + exact h x hx + +lemma norm_cfc_le_iff (f : 𝕜 → 𝕜) (a : A) {c : ℝ} (hc : 0 ≤ c) + (hf : ContinuousOn f (σ 𝕜 a) := by cfc_cont_tac) + (ha : p a := by cfc_tac) : ‖cfc f a‖ ≤ c ↔ ∀ x ∈ σ 𝕜 a, ‖f x‖ ≤ c := + ⟨fun h _ hx ↦ norm_apply_le_norm_cfc f a hx hf ha |>.trans h, norm_cfc_le hc⟩ + +open NNReal + +lemma nnnorm_cfc_le {f : 𝕜 → 𝕜} {a : A} (c : ℝ≥0) (h : ∀ x ∈ σ 𝕜 a, ‖f x‖₊ ≤ c) : + ‖cfc f a‖₊ ≤ c := + norm_cfc_le c.2 h + +lemma nnnorm_cfc_le_iff (f : 𝕜 → 𝕜) (a : A) (c : ℝ≥0) + (hf : ContinuousOn f (σ 𝕜 a) := by cfc_cont_tac) + (ha : p a := by cfc_tac) : ‖cfc f a‖₊ ≤ c ↔ ∀ x ∈ σ 𝕜 a, ‖f x‖₊ ≤ c := + norm_cfc_le_iff f a c.2 + +end NormedRing + +namespace SpectrumRestricts + +variable {R S A : Type*} {p q : A → Prop} +variable [Semifield R] [StarRing R] [MetricSpace R] [TopologicalSemiring R] [ContinuousStar R] +variable [Semifield S] [StarRing S] [MetricSpace S] [TopologicalSemiring S] [ContinuousStar S] +variable [Ring A] [StarRing A] [Algebra S A] +variable [Algebra R S] [Algebra R A] [IsScalarTower R S A] [StarModule R S] [ContinuousSMul R S] +variable [MetricSpace A] [IsometricContinuousFunctionalCalculus S A q] +variable [CompleteSpace R] [UniqueContinuousFunctionalCalculus R A] + +open scoped ContinuousFunctionalCalculus in +protected theorem isometric_cfc (f : C(S, R)) (halg : Isometry (algebraMap R S)) (h0 : p 0) + (h : ∀ a, p a ↔ q a ∧ SpectrumRestricts a f) : + IsometricContinuousFunctionalCalculus R A p where + toContinuousFunctionalCalculus := SpectrumRestricts.cfc f halg.isUniformEmbedding h0 h + isometric a ha := by + obtain ⟨ha', haf⟩ := h a |>.mp ha + have _inst (a : A) : CompactSpace (σ R a) := by + rw [← isCompact_iff_compactSpace, ← spectrum.preimage_algebraMap S] + exact halg.isClosedEmbedding.isCompact_preimage <| + ContinuousFunctionalCalculus.isCompact_spectrum a + have := SpectrumRestricts.cfc f halg.isUniformEmbedding h0 h + rw [cfcHom_eq_restrict f halg.isUniformEmbedding ha ha' haf] + refine .of_dist_eq fun g₁ g₂ ↦ ?_ + simp only [starAlgHom_apply, isometry_cfcHom a ha' |>.dist_eq] + refine le_antisymm ?_ ?_ + all_goals refine ContinuousMap.dist_le dist_nonneg |>.mpr fun x ↦ ?_ + · simpa [halg.dist_eq] using ContinuousMap.dist_apply_le_dist _ + · let x' : σ S a := Subtype.map (algebraMap R S) (fun _ ↦ spectrum.algebraMap_mem S) x + apply le_of_eq_of_le ?_ <| ContinuousMap.dist_apply_le_dist x' + simp only [ContinuousMap.comp_apply, ContinuousMap.coe_mk, StarAlgHom.ofId_apply, + halg.dist_eq, x'] + congr! + all_goals ext; exact haf.left_inv _ |>.symm + +end SpectrumRestricts + +end Unital + +/-! ### Isometric continuous functional calculus for non-unital algebras -/ + +section NonUnital + +/-- An extension of the `NonUnitalContinuousFunctionalCalculus` requiring that `cfcₙHom` is an +isometry. -/ +class NonUnitalIsometricContinuousFunctionalCalculus (R A : Type*) (p : outParam (A → Prop)) + [CommSemiring R] [Nontrivial R] [StarRing R] [MetricSpace R] [TopologicalSemiring R] + [ContinuousStar R] [NonUnitalRing A] [StarRing A] [MetricSpace A] [Module R A] + [IsScalarTower R A A] [SMulCommClass R A A] + extends NonUnitalContinuousFunctionalCalculus R p : Prop where + isometric (a : A) (ha : p a) : Isometry (cfcₙHom ha (R := R)) + +section MetricSpace + +variable {R A : Type*} {p : outParam (A → Prop)} +variable [CommSemiring R] [Nontrivial R] [StarRing R] [MetricSpace R] [TopologicalSemiring R] +variable [ContinuousStar R] +variable [NonUnitalRing A] [StarRing A] [MetricSpace A] [Module R A] +variable [IsScalarTower R A A] [SMulCommClass R A A] + +open scoped NonUnitalContinuousFunctionalCalculus + +variable [NonUnitalIsometricContinuousFunctionalCalculus R A p] + +lemma isometry_cfcₙHom (a : A) (ha : p a := by cfc_tac) : + Isometry (cfcₙHom (show p a from ha) (R := R)) := + NonUnitalIsometricContinuousFunctionalCalculus.isometric a ha + +end MetricSpace + +section NormedRing + +variable {𝕜 A : Type*} {p : outParam (A → Prop)} +variable [RCLike 𝕜] [NonUnitalNormedRing A] [StarRing A] [NormedSpace 𝕜 A] [IsScalarTower 𝕜 A A] +variable [SMulCommClass 𝕜 A A] +variable [NonUnitalIsometricContinuousFunctionalCalculus 𝕜 A p] + +open NonUnitalIsometricContinuousFunctionalCalculus +open scoped ContinuousMapZero NonUnitalContinuousFunctionalCalculus + +lemma norm_cfcₙHom (a : A) (f : C(σₙ 𝕜 a, 𝕜)₀) (ha : p a := by cfc_tac) : + ‖cfcₙHom (show p a from ha) f‖ = ‖f‖ := by + refine isometry_cfcₙHom a |>.norm_map_of_map_zero (map_zero _) f + +lemma nnnorm_cfcₙHom (a : A) (f : C(σₙ 𝕜 a, 𝕜)₀) (ha : p a := by cfc_tac) : + ‖cfcₙHom (show p a from ha) f‖₊ = ‖f‖₊ := + Subtype.ext <| norm_cfcₙHom a f ha + +lemma IsGreatest.norm_cfcₙ (f : 𝕜 → 𝕜) (a : A) + (hf : ContinuousOn f (σₙ 𝕜 a) := by cfc_cont_tac) (hf₀ : f 0 = 0 := by cfc_zero_tac) + (ha : p a := by cfc_tac) : IsGreatest ((fun x ↦ ‖f x‖) '' σₙ 𝕜 a) ‖cfcₙ f a‖ := by + obtain ⟨x, hx⟩ := NonUnitalContinuousFunctionalCalculus.isCompact_quasispectrum a + |>.image_of_continuousOn hf.norm |>.exists_isGreatest <| + (quasispectrum.nonempty 𝕜 a).image _ + obtain ⟨x, hx', rfl⟩ := hx.1 + convert hx + rw [cfcₙ_apply f a, norm_cfcₙHom a _] + apply le_antisymm + · apply ContinuousMap.norm_le _ (norm_nonneg _) |>.mpr + rintro ⟨y, hy⟩ + exact hx.2 ⟨y, hy, rfl⟩ + · exact le_trans (by simp) <| ContinuousMap.norm_coe_le_norm _ (⟨x, hx'⟩ : σₙ 𝕜 a) + +lemma IsGreatest.nnnorm_cfcₙ (f : 𝕜 → 𝕜) (a : A) + (hf : ContinuousOn f (σₙ 𝕜 a) := by cfc_cont_tac) (hf₀ : f 0 = 0 := by cfc_zero_tac) + (ha : p a := by cfc_tac) : IsGreatest ((fun x ↦ ‖f x‖₊) '' σₙ 𝕜 a) ‖cfcₙ f a‖₊ := by + convert Real.toNNReal_mono.map_isGreatest (.norm_cfcₙ f a) + all_goals simp [Set.image_image, norm_toNNReal] + +lemma norm_apply_le_norm_cfcₙ (f : 𝕜 → 𝕜) (a : A) ⦃x : 𝕜⦄ (hx : x ∈ σₙ 𝕜 a) + (hf : ContinuousOn f (σₙ 𝕜 a) := by cfc_cont_tac) (hf₀ : f 0 = 0 := by cfc_zero_tac) + (ha : p a := by cfc_tac) : ‖f x‖ ≤ ‖cfcₙ f a‖ := + IsGreatest.norm_cfcₙ f a hf hf₀ ha |>.2 ⟨x, hx, rfl⟩ + +lemma nnnorm_apply_le_nnnorm_cfcₙ (f : 𝕜 → 𝕜) (a : A) ⦃x : 𝕜⦄ (hx : x ∈ σₙ 𝕜 a) + (hf : ContinuousOn f (σₙ 𝕜 a) := by cfc_cont_tac) (hf₀ : f 0 = 0 := by cfc_zero_tac) + (ha : p a := by cfc_tac) : ‖f x‖₊ ≤ ‖cfcₙ f a‖₊ := + IsGreatest.nnnorm_cfcₙ f a hf hf₀ ha |>.2 ⟨x, hx, rfl⟩ + +lemma norm_cfcₙ_le {f : 𝕜 → 𝕜} {a : A} {c : ℝ} (h : ∀ x ∈ σₙ 𝕜 a, ‖f x‖ ≤ c) : + ‖cfcₙ f a‖ ≤ c := by + refine cfcₙ_cases (‖·‖ ≤ c) a f ?_ fun hf hf0 ha ↦ ?_ + · simpa using (norm_nonneg _).trans <| h 0 (quasispectrum.zero_mem 𝕜 a) + · simp only [← cfcₙ_apply f a, isLUB_le_iff (IsGreatest.norm_cfcₙ f a hf hf0 ha |>.isLUB)] + rintro - ⟨x, hx, rfl⟩ + exact h x hx + +lemma norm_cfcₙ_le_iff (f : 𝕜 → 𝕜) (a : A) (c : ℝ) + (hf : ContinuousOn f (σₙ 𝕜 a) := by cfc_cont_tac) (hf₀ : f 0 = 0 := by cfc_zero_tac) + (ha : p a := by cfc_tac) : ‖cfcₙ f a‖ ≤ c ↔ ∀ x ∈ σₙ 𝕜 a, ‖f x‖ ≤ c := + ⟨fun h _ hx ↦ norm_apply_le_norm_cfcₙ f a hx hf hf₀ ha |>.trans h, norm_cfcₙ_le⟩ + +open NNReal + +lemma nnnorm_cfcₙ_le {f : 𝕜 → 𝕜} {a : A} {c : ℝ≥0} (h : ∀ x ∈ σₙ 𝕜 a, ‖f x‖₊ ≤ c) : + ‖cfcₙ f a‖₊ ≤ c := + norm_cfcₙ_le h + +lemma nnnorm_cfcₙ_le_iff (f : 𝕜 → 𝕜) (a : A) (c : ℝ≥0) + (hf : ContinuousOn f (σₙ 𝕜 a) := by cfc_cont_tac) (hf₀ : f 0 = 0 := by cfc_zero_tac) + (ha : p a := by cfc_tac) : ‖cfcₙ f a‖₊ ≤ c ↔ ∀ x ∈ σₙ 𝕜 a, ‖f x‖₊ ≤ c := + norm_cfcₙ_le_iff f a c.1 hf hf₀ ha + +end NormedRing + +namespace QuasispectrumRestricts + +open NonUnitalIsometricContinuousFunctionalCalculus + +variable {R S A : Type*} {p q : A → Prop} +variable [Semifield R] [StarRing R] [MetricSpace R] [TopologicalSemiring R] [ContinuousStar R] +variable [Field S] [StarRing S] [MetricSpace S] [TopologicalRing S] [ContinuousStar S] +variable [NonUnitalRing A] [StarRing A] [Module S A] [IsScalarTower S A A] +variable [SMulCommClass S A A] +variable [Algebra R S] [Module R A] [IsScalarTower R S A] [StarModule R S] [ContinuousSMul R S] +variable [IsScalarTower R A A] [SMulCommClass R A A] +variable [MetricSpace A] [NonUnitalIsometricContinuousFunctionalCalculus S A q] +variable [CompleteSpace R] [UniqueNonUnitalContinuousFunctionalCalculus R A] + +open scoped NonUnitalContinuousFunctionalCalculus in +protected theorem isometric_cfc (f : C(S, R)) (halg : Isometry (algebraMap R S)) (h0 : p 0) + (h : ∀ a, p a ↔ q a ∧ QuasispectrumRestricts a f) : + NonUnitalIsometricContinuousFunctionalCalculus R A p where + toNonUnitalContinuousFunctionalCalculus := QuasispectrumRestricts.cfc f + halg.isUniformEmbedding h0 h + isometric a ha := by + obtain ⟨ha', haf⟩ := h a |>.mp ha + have _inst (a : A) : CompactSpace (σₙ R a) := by + rw [← isCompact_iff_compactSpace, ← quasispectrum.preimage_algebraMap S] + exact halg.isClosedEmbedding.isCompact_preimage <| + NonUnitalContinuousFunctionalCalculus.isCompact_quasispectrum a + have := QuasispectrumRestricts.cfc f halg.isUniformEmbedding h0 h + rw [cfcₙHom_eq_restrict f halg.isUniformEmbedding ha ha' haf] + refine .of_dist_eq fun g₁ g₂ ↦ ?_ + simp only [nonUnitalStarAlgHom_apply, isometry_cfcₙHom a ha' |>.dist_eq] + refine le_antisymm ?_ ?_ + all_goals refine ContinuousMap.dist_le dist_nonneg |>.mpr fun x ↦ ?_ + · simpa [halg.dist_eq] using ContinuousMap.dist_apply_le_dist _ + · let x' : σₙ S a := Subtype.map (algebraMap R S) (fun _ ↦ quasispectrum.algebraMap_mem S) x + apply le_of_eq_of_le ?_ <| ContinuousMap.dist_apply_le_dist x' + simp only [ContinuousMap.coe_coe, ContinuousMapZero.comp_apply, ContinuousMapZero.coe_mk, + ContinuousMap.coe_mk, StarAlgHom.ofId_apply, halg.dist_eq, x'] + congr! 2 + all_goals ext; exact haf.left_inv _ |>.symm + +end QuasispectrumRestricts + +end NonUnital + +/-! ### Instances of isometric continuous functional calculi -/ + +section Instances + +section Unital + +section Complex + +variable {A : Type*} [CStarAlgebra A] + +instance IsStarNormal.instIsometricContinuousFunctionalCalculus : + IsometricContinuousFunctionalCalculus ℂ A IsStarNormal where + isometric a ha := by + rw [cfcHom_eq_of_isStarNormal] + exact isometry_subtype_coe.comp <| StarAlgEquiv.isometry (continuousFunctionalCalculus a) + +instance IsSelfAdjoint.instIsometricContinuousFunctionalCalculus : + IsometricContinuousFunctionalCalculus ℝ A IsSelfAdjoint := + SpectrumRestricts.isometric_cfc Complex.reCLM Complex.isometry_ofReal (.zero _) + fun _ ↦ isSelfAdjoint_iff_isStarNormal_and_spectrumRestricts + +end Complex + +section NNReal + +variable {A : Type*} [NormedRing A] [PartialOrder A] [StarRing A] [StarOrderedRing A] +variable [NormedAlgebra ℝ A] [IsometricContinuousFunctionalCalculus ℝ A IsSelfAdjoint] +variable [NonnegSpectrumClass ℝ A] [UniqueContinuousFunctionalCalculus ℝ A] + +open NNReal in +instance Nonneg.instIsometricContinuousFunctionalCalculus : + IsometricContinuousFunctionalCalculus ℝ≥0 A (0 ≤ ·) := + SpectrumRestricts.isometric_cfc (q := IsSelfAdjoint) ContinuousMap.realToNNReal + isometry_subtype_coe le_rfl (fun _ ↦ nonneg_iff_isSelfAdjoint_and_spectrumRestricts) + +end NNReal + +end Unital + +section NonUnital + +section Complex + +variable {A : Type*} [NonUnitalCStarAlgebra A] + +open Unitization + + +open ContinuousMapZero in +instance IsStarNormal.instNonUnitalIsometricContinuousFunctionalCalculus : + NonUnitalIsometricContinuousFunctionalCalculus ℂ A IsStarNormal where + isometric a ha := by + refine AddMonoidHomClass.isometry_of_norm _ fun f ↦ ?_ + rw [← norm_inr (𝕜 := ℂ), ← inrNonUnitalStarAlgHom_apply, ← NonUnitalStarAlgHom.comp_apply, + inr_comp_cfcₙHom_eq_cfcₙAux a, cfcₙAux] + simp only [NonUnitalStarAlgHom.comp_assoc, NonUnitalStarAlgHom.comp_apply, + toContinuousMapHom_apply, NonUnitalStarAlgHom.coe_coe] + rw [norm_cfcHom (a : Unitization ℂ A), StarAlgEquiv.norm_map] + rfl + +instance IsSelfAdjoint.instNonUnitalIsometricContinuousFunctionalCalculus : + NonUnitalIsometricContinuousFunctionalCalculus ℝ A IsSelfAdjoint := + QuasispectrumRestricts.isometric_cfc Complex.reCLM Complex.isometry_ofReal (.zero _) + fun _ ↦ isSelfAdjoint_iff_isStarNormal_and_quasispectrumRestricts + +end Complex + +section NNReal + +variable {A : Type*} [NonUnitalNormedRing A] [PartialOrder A] [StarRing A] [StarOrderedRing A] +variable [NormedSpace ℝ A] [IsScalarTower ℝ A A] [SMulCommClass ℝ A A] +variable [NonUnitalIsometricContinuousFunctionalCalculus ℝ A IsSelfAdjoint] +variable [NonnegSpectrumClass ℝ A] [UniqueNonUnitalContinuousFunctionalCalculus ℝ A] + +open NNReal in +instance Nonneg.instNonUnitalIsometricContinuousFunctionalCalculus : + NonUnitalIsometricContinuousFunctionalCalculus ℝ≥0 A (0 ≤ ·) := + QuasispectrumRestricts.isometric_cfc (q := IsSelfAdjoint) ContinuousMap.realToNNReal + isometry_subtype_coe le_rfl (fun _ ↦ nonneg_iff_isSelfAdjoint_and_quasispectrumRestricts) + +end NNReal + +end NonUnital + +end Instances + +/-! ### Properties specific to `ℝ≥0` -/ + +section NNReal + +open NNReal + +section Unital + +variable {A : Type*} [NormedRing A] [StarRing A] [NormedAlgebra ℝ A] [PartialOrder A] +variable [StarOrderedRing A] [IsometricContinuousFunctionalCalculus ℝ A IsSelfAdjoint] +variable [NonnegSpectrumClass ℝ A] [UniqueContinuousFunctionalCalculus ℝ A] + +lemma IsGreatest.nnnorm_cfc_nnreal [Nontrivial A] (f : ℝ≥0 → ℝ≥0) (a : A) + (hf : ContinuousOn f (σ ℝ≥0 a) := by cfc_cont_tac) (ha : 0 ≤ a := by cfc_tac) : + IsGreatest (f '' σ ℝ≥0 a) ‖cfc f a‖₊ := by + rw [cfc_nnreal_eq_real] + obtain ⟨-, ha'⟩ := nonneg_iff_isSelfAdjoint_and_spectrumRestricts.mp ha + convert IsGreatest.nnnorm_cfc (fun x : ℝ ↦ (f x.toNNReal : ℝ)) a ?hf_cont + case hf_cont => exact continuous_subtype_val.comp_continuousOn <| + ContinuousOn.comp ‹_› continuous_real_toNNReal.continuousOn <| ha'.image ▸ Set.mapsTo_image .. + ext x + constructor + all_goals rintro ⟨x, hx, rfl⟩ + · exact ⟨x, spectrum.algebraMap_mem ℝ hx, by simp⟩ + · exact ⟨x.toNNReal, ha'.apply_mem hx, by simp⟩ + +lemma apply_le_nnnorm_cfc_nnreal (f : ℝ≥0 → ℝ≥0) (a : A) ⦃x : ℝ≥0⦄ (hx : x ∈ σ ℝ≥0 a) + (hf : ContinuousOn f (σ ℝ≥0 a) := by cfc_cont_tac) (ha : 0 ≤ a := by cfc_tac) : + f x ≤ ‖cfc f a‖₊ := by + revert hx + nontriviality A + exact (IsGreatest.nnnorm_cfc_nnreal f a hf ha |>.2 ⟨x, ·, rfl⟩) + +lemma nnnorm_cfc_nnreal_le {f : ℝ≥0 → ℝ≥0} {a : A} {c : ℝ≥0} (h : ∀ x ∈ σ ℝ≥0 a, f x ≤ c) : + ‖cfc f a‖₊ ≤ c := by + obtain (_ | _) := subsingleton_or_nontrivial A + · rw [Subsingleton.elim (cfc f a) 0] + simp + · refine cfc_cases (‖·‖₊ ≤ c) a f (by simp) fun hf ha ↦ ?_ + simp only [← cfc_apply f a, isLUB_le_iff (IsGreatest.nnnorm_cfc_nnreal f a hf ha |>.isLUB)] + rintro - ⟨x, hx, rfl⟩ + exact h x hx + +lemma nnnorm_cfc_nnreal_le_iff (f : ℝ≥0 → ℝ≥0) (a : A) (c : ℝ≥0) + (hf : ContinuousOn f (σ ℝ≥0 a) := by cfc_cont_tac) + (ha : 0 ≤ a := by cfc_tac) : ‖cfc f a‖₊ ≤ c ↔ ∀ x ∈ σ ℝ≥0 a, f x ≤ c := + ⟨fun h _ hx ↦ apply_le_nnnorm_cfc_nnreal f a hx hf ha |>.trans h, nnnorm_cfc_nnreal_le⟩ + +end Unital + +section NonUnital + +variable {A : Type*} [NonUnitalNormedRing A] [StarRing A] [NormedSpace ℝ A] +variable [IsScalarTower ℝ A A] [SMulCommClass ℝ A A] [PartialOrder A] +variable [StarOrderedRing A] [NonUnitalIsometricContinuousFunctionalCalculus ℝ A IsSelfAdjoint] +variable [NonnegSpectrumClass ℝ A] [UniqueNonUnitalContinuousFunctionalCalculus ℝ A] + +lemma IsGreatest.nnnorm_cfcₙ_nnreal (f : ℝ≥0 → ℝ≥0) (a : A) + (hf : ContinuousOn f (σₙ ℝ≥0 a) := by cfc_cont_tac) (hf0 : f 0 = 0 := by cfc_zero_tac) + (ha : 0 ≤ a := by cfc_tac) : IsGreatest (f '' σₙ ℝ≥0 a) ‖cfcₙ f a‖₊ := by + rw [cfcₙ_nnreal_eq_real] + obtain ⟨-, ha'⟩ := nonneg_iff_isSelfAdjoint_and_quasispectrumRestricts.mp ha + convert IsGreatest.nnnorm_cfcₙ (fun x : ℝ ↦ (f x.toNNReal : ℝ)) a ?hf_cont (by simpa) + case hf_cont => exact continuous_subtype_val.comp_continuousOn <| + ContinuousOn.comp ‹_› continuous_real_toNNReal.continuousOn <| ha'.image ▸ Set.mapsTo_image .. + ext x + constructor + all_goals rintro ⟨x, hx, rfl⟩ + · exact ⟨x, quasispectrum.algebraMap_mem ℝ hx, by simp⟩ + · exact ⟨x.toNNReal, ha'.apply_mem hx, by simp⟩ + +lemma apply_le_nnnorm_cfcₙ_nnreal (f : ℝ≥0 → ℝ≥0) (a : A) ⦃x : ℝ≥0⦄ (hx : x ∈ σₙ ℝ≥0 a) + (hf : ContinuousOn f (σₙ ℝ≥0 a) := by cfc_cont_tac) (hf0 : f 0 = 0 := by cfc_zero_tac) + (ha : 0 ≤ a := by cfc_tac) : f x ≤ ‖cfcₙ f a‖₊ := by + revert hx + exact (IsGreatest.nnnorm_cfcₙ_nnreal f a hf hf0 ha |>.2 ⟨x, ·, rfl⟩) + +lemma nnnorm_cfcₙ_nnreal_le {f : ℝ≥0 → ℝ≥0} {a : A} {c : ℝ≥0} (h : ∀ x ∈ σₙ ℝ≥0 a, f x ≤ c) : + ‖cfcₙ f a‖₊ ≤ c := by + refine cfcₙ_cases (‖·‖₊ ≤ c) a f (by simp) fun hf hf0 ha ↦ ?_ + simp only [← cfcₙ_apply f a, isLUB_le_iff (IsGreatest.nnnorm_cfcₙ_nnreal f a hf hf0 ha |>.isLUB)] + rintro - ⟨x, hx, rfl⟩ + exact h x hx + +lemma nnnorm_cfcₙ_nnreal_le_iff (f : ℝ≥0 → ℝ≥0) (a : A) (c : ℝ≥0) + (hf : ContinuousOn f (σₙ ℝ≥0 a) := by cfc_cont_tac) (hf₀ : f 0 = 0 := by cfc_zero_tac) + (ha : 0 ≤ a := by cfc_tac) : ‖cfcₙ f a‖₊ ≤ c ↔ ∀ x ∈ σₙ ℝ≥0 a, f x ≤ c := + ⟨fun h _ hx ↦ apply_le_nnnorm_cfcₙ_nnreal f a hx hf hf₀ ha |>.trans h, nnnorm_cfcₙ_nnreal_le⟩ + +end NonUnital + +end NNReal diff --git a/Mathlib/Analysis/Normed/Algebra/Spectrum.lean b/Mathlib/Analysis/Normed/Algebra/Spectrum.lean index 942cd2eeae3a2..6c924b2091498 100644 --- a/Mathlib/Analysis/Normed/Algebra/Spectrum.lean +++ b/Mathlib/Analysis/Normed/Algebra/Spectrum.lean @@ -158,6 +158,23 @@ instance _root_.quasispectrum.instCompactSpaceNNReal [NormedSpace ℝ B] [IsScal end QuasispectrumCompact +section NNReal + +open NNReal + +variable {A : Type*} [NormedRing A] [NormedAlgebra ℝ A] [CompleteSpace A] [NormOneClass A] + +theorem le_nnnorm_of_mem {a : A} {r : ℝ≥0} (hr : r ∈ spectrum ℝ≥0 a) : + r ≤ ‖a‖₊ := calc + r ≤ ‖(r : ℝ)‖ := Real.le_norm_self _ + _ ≤ ‖a‖ := norm_le_norm_of_mem hr + +theorem coe_le_norm_of_mem {a : A} {r : ℝ≥0} (hr : r ∈ spectrum ℝ≥0 a) : + r ≤ ‖a‖ := + coe_mono <| le_nnnorm_of_mem hr + +end NNReal + theorem spectralRadius_le_nnnorm [NormOneClass A] (a : A) : spectralRadius 𝕜 a ≤ ‖a‖₊ := by refine iSup₂_le fun k hk => ?_ exact mod_cast norm_le_norm_of_mem hk diff --git a/Mathlib/Analysis/Normed/Algebra/Unitization.lean b/Mathlib/Analysis/Normed/Algebra/Unitization.lean index de8b39cdbddce..8212d68f7249e 100644 --- a/Mathlib/Analysis/Normed/Algebra/Unitization.lean +++ b/Mathlib/Analysis/Normed/Algebra/Unitization.lean @@ -256,6 +256,10 @@ lemma nnnorm_inr (a : A) : ‖(a : Unitization 𝕜 A)‖₊ = ‖a‖₊ := lemma isometry_inr : Isometry ((↑) : A → Unitization 𝕜 A) := AddMonoidHomClass.isometry_of_norm (inrNonUnitalAlgHom 𝕜 A) norm_inr +@[fun_prop] +theorem continuous_inr : Continuous (inr : A → Unitization 𝕜 A) := + isometry_inr.continuous + lemma dist_inr (a b : A) : dist (a : Unitization 𝕜 A) (b : Unitization 𝕜 A) = dist a b := isometry_inr.dist_eq a b From 2061d034b2e163f02fdbe7c36548662355406fee Mon Sep 17 00:00:00 2001 From: blizzard_inc Date: Sun, 27 Oct 2024 22:45:06 +0000 Subject: [PATCH 490/505] fix(Util/FormatTable): escape bar characters in formatting tables for markdown (#18273) This PR changes the `formatTable` function to make sure vertical bar (`|`) characters get escaped when formatting strings into markdown tables. Co-authored-by: Michael Rothgang --- Mathlib/Util/FormatTable.lean | 12 ++++++++---- test/format_table.lean | 17 +++++++++++++++++ 2 files changed, 25 insertions(+), 4 deletions(-) create mode 100644 test/format_table.lean diff --git a/Mathlib/Util/FormatTable.lean b/Mathlib/Util/FormatTable.lean index 036f4dfd32366..57d5b6dbf2afb 100644 --- a/Mathlib/Util/FormatTable.lean +++ b/Mathlib/Util/FormatTable.lean @@ -38,14 +38,18 @@ def formatTable (headers : Array String) (table : Array (Array String)) String := Id.run do -- If no alignments are provided, default to left alignment for all columns. let alignments := alignments.getD (Array.mkArray headers.size Alignment.left) + -- Escape all vertical bar characters inside a table cell, + -- otherwise these could get interpreted as starting a new row or column. + let escapedHeaders := headers.map (fun header => header.replace "|" "\\|") + let escapedTable := table.map (fun row => row.map (fun cell => cell.replace "|" "\\|")) -- Compute the maximum width of each column. - let mut widths := headers.map (·.length) - for row in table do + let mut widths := escapedHeaders.map (·.length) + for row in escapedTable do for i in [0:widths.size] do widths := widths.set! i (max widths[i]! ((row[i]?.map (·.length)).getD 0)) -- Pad each cell with spaces to match the column width. - let paddedHeaders := headers.mapIdx fun i h => h.rightpad widths[i]! - let paddedTable := table.map fun row => row.mapIdx fun i cell => + let paddedHeaders := escapedHeaders.mapIdx fun i h => h.rightpad widths[i]! + let paddedTable := escapedTable.map fun row => row.mapIdx fun i cell => cell.justify alignments[i]! widths[i]! -- Construct the lines of the table let headerLine := "| " ++ String.intercalate " | " (paddedHeaders.toList) ++ " |" diff --git a/test/format_table.lean b/test/format_table.lean new file mode 100644 index 0000000000000..296df3992f7ec --- /dev/null +++ b/test/format_table.lean @@ -0,0 +1,17 @@ +import Mathlib.Util.FormatTable + +/-- +info: +| first | second header | third \| header | +| :-------------- | :-------------: | --------------: | +| item number one | item two | item c | +| the fourth item | the fourth item | escape \| this | +| align left | align center | align right | +-/ +#guard_msgs in +run_cmd Lean.logInfo ("\n" ++ formatTable + #["first", "second header", "third | header"] + #[#["item number one", "item two", "item c"], + #["the fourth item", "the fourth item", "escape | this"], + #["align left", "align center", "align right"]] + (.some #[.left, .center, .right])) From 91530d46942aa5d14a89d26826f87096a8c536e9 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Mon, 28 Oct 2024 00:34:35 +0000 Subject: [PATCH 491/505] feat: some trivialities about Nat.primeCounting and Finset.card/count (#17037) Also move some general results out of `PrimeCounting.lean`. --- Mathlib.lean | 1 + Mathlib/Data/Finset/Card.lean | 20 ++++++++++++++++---- Mathlib/Data/Nat/Count.lean | 20 ++++++++++++++------ Mathlib/Data/Nat/Nth.lean | 3 +++ Mathlib/Data/Nat/Prime/Nth.lean | 23 +++++++++++++++++++++++ Mathlib/NumberTheory/PrimeCounting.lean | 23 ++++++++++++++++++----- 6 files changed, 75 insertions(+), 15 deletions(-) create mode 100644 Mathlib/Data/Nat/Prime/Nth.lean diff --git a/Mathlib.lean b/Mathlib.lean index cc6e5cead888e..e90ee6329b3f6 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2555,6 +2555,7 @@ import Mathlib.Data.Nat.Prime.Defs import Mathlib.Data.Nat.Prime.Factorial import Mathlib.Data.Nat.Prime.Infinite import Mathlib.Data.Nat.Prime.Int +import Mathlib.Data.Nat.Prime.Nth import Mathlib.Data.Nat.Prime.Pow import Mathlib.Data.Nat.PrimeFin import Mathlib.Data.Nat.Set diff --git a/Mathlib/Data/Finset/Card.lean b/Mathlib/Data/Finset/Card.lean index 67ac01fae4261..b08d5c973026f 100644 --- a/Mathlib/Data/Finset/Card.lean +++ b/Mathlib/Data/Finset/Card.lean @@ -276,19 +276,31 @@ theorem card_filter_le (s : Finset α) (p : α → Prop) [DecidablePred p] : theorem eq_of_subset_of_card_le {s t : Finset α} (h : s ⊆ t) (h₂ : #t ≤ #s) : s = t := eq_of_veq <| Multiset.eq_of_le_of_card_le (val_le_iff.mpr h) h₂ +theorem eq_iff_card_le_of_subset (hst : s ⊆ t) : #t ≤ #s ↔ s = t := + ⟨eq_of_subset_of_card_le hst, (ge_of_eq <| congr_arg _ ·)⟩ + theorem eq_of_superset_of_card_ge (hst : s ⊆ t) (hts : #t ≤ #s) : t = s := (eq_of_subset_of_card_le hst hts).symm +theorem eq_iff_card_ge_of_superset (hst : s ⊆ t) : #t ≤ #s ↔ t = s := + (eq_iff_card_le_of_subset hst).trans eq_comm + theorem subset_iff_eq_of_card_le (h : #t ≤ #s) : s ⊆ t ↔ s = t := ⟨fun hst => eq_of_subset_of_card_le hst h, Eq.subset'⟩ theorem map_eq_of_subset {f : α ↪ α} (hs : s.map f ⊆ s) : s.map f = s := eq_of_subset_of_card_le hs (card_map _).ge -theorem filter_card_eq {p : α → Prop} [DecidablePred p] (h : #(s.filter p) = #s) (x : α) - (hx : x ∈ s) : p x := by - rw [← eq_of_subset_of_card_le (s.filter_subset p) h.ge, mem_filter] at hx - exact hx.2 +theorem card_filter_eq_iff {p : α → Prop} [DecidablePred p] : + #(s.filter p) = #s ↔ ∀ x ∈ s, p x := by + rw [(card_filter_le s p).eq_iff_not_lt, not_lt, eq_iff_card_le_of_subset (filter_subset p s), + filter_eq_self] + +alias ⟨filter_card_eq, _⟩ := card_filter_eq_iff + +theorem card_filter_eq_zero_iff {p : α → Prop} [DecidablePred p] : + #(s.filter p) = 0 ↔ ∀ x ∈ s, ¬ p x := by + rw [card_eq_zero, filter_eq_empty_iff] nonrec lemma card_lt_card (h : s ⊂ t) : #s < #t := card_lt_card <| val_lt_iff.2 h diff --git a/Mathlib/Data/Nat/Count.lean b/Mathlib/Data/Nat/Count.lean index 3cc54f507cd9c..4d8294ce6bceb 100644 --- a/Mathlib/Data/Nat/Count.lean +++ b/Mathlib/Data/Nat/Count.lean @@ -54,6 +54,10 @@ theorem count_eq_card_fintype (n : ℕ) : count p n = Fintype.card { k : ℕ // rw [count_eq_card_filter_range, ← Fintype.card_ofFinset, ← CountSet.fintype] rfl +theorem count_le {n : ℕ} : count p n ≤ n := by + rw [count_eq_card_filter_range] + exact (card_filter_le _ _).trans_eq (card_range _) + theorem count_succ (n : ℕ) : count p (n + 1) = count p n + if p n then 1 else 0 := by split_ifs with h <;> simp [count, List.range_succ, h] @@ -121,15 +125,19 @@ theorem count_le_card (hp : (setOf p).Finite) (n : ℕ) : count p n ≤ #hp.toFi theorem count_lt_card {n : ℕ} (hp : (setOf p).Finite) (hpn : p n) : count p n < #hp.toFinset := (count_lt_count_succ_iff.2 hpn).trans_le (count_le_card hp _) -theorem count_of_forall {n : ℕ} (hp : ∀ n' < n, p n') : count p n = n := by - rw [count_eq_card_filter_range, filter_true_of_mem, card_range] - · simpa only [Finset.mem_range] +theorem count_iff_forall {n : ℕ} : count p n = n ↔ ∀ n' < n, p n' := by + simpa [count_eq_card_filter_range, card_range, mem_range] using + card_filter_eq_iff (p := p) (s := range n) + +alias ⟨_, count_of_forall⟩ := count_iff_forall @[simp] theorem count_true (n : ℕ) : count (fun _ ↦ True) n = n := count_of_forall fun _ _ ↦ trivial -theorem count_of_forall_not {n : ℕ} (hp : ∀ n' < n, ¬p n') : count p n = 0 := by - rw [count_eq_card_filter_range, filter_false_of_mem, card_empty] - · simpa only [Finset.mem_range] +theorem count_iff_forall_not {n : ℕ} : count p n = 0 ↔ ∀ m < n, ¬p m := by + simpa [count_eq_card_filter_range, mem_range] using + card_filter_eq_zero_iff (p := p) (s := range n) + +alias ⟨_, count_of_forall_not⟩ := count_iff_forall_not @[simp] theorem count_false (n : ℕ) : count (fun _ ↦ False) n = 0 := count_of_forall_not fun _ _ ↦ id diff --git a/Mathlib/Data/Nat/Nth.lean b/Mathlib/Data/Nat/Nth.lean index abe7621d093a2..fb3e78daa3035 100644 --- a/Mathlib/Data/Nat/Nth.lean +++ b/Mathlib/Data/Nat/Nth.lean @@ -323,6 +323,9 @@ theorem nth_lt_of_lt_count {n k : ℕ} (h : k < count p n) : nth p k < n := by theorem le_nth_of_count_le {n k : ℕ} (h : n ≤ nth p k) : count p n ≤ k := not_lt.1 fun hlt => h.not_lt <| nth_lt_of_lt_count hlt +protected theorem count_eq_zero (h : ∃ n, p n) {n : ℕ} : count p n = 0 ↔ n ≤ nth p 0 := by + rw [nth_zero_of_exists h, le_find_iff h, Nat.count_iff_forall_not] + variable (p) theorem nth_count_eq_sInf (n : ℕ) : nth p (count p n) = sInf {i : ℕ | p i ∧ n ≤ i} := by diff --git a/Mathlib/Data/Nat/Prime/Nth.lean b/Mathlib/Data/Nat/Prime/Nth.lean new file mode 100644 index 0000000000000..bc59dbf9f528f --- /dev/null +++ b/Mathlib/Data/Nat/Prime/Nth.lean @@ -0,0 +1,23 @@ +/- +Copyright (c) 2024 Ralf Stephan. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Ralf Stephan +-/ +import Mathlib.Data.Nat.Prime.Defs +import Mathlib.Data.Nat.Nth + +/-! +# The Nth primes +-/ + +namespace Nat + +@[simp] +theorem nth_prime_zero_eq_two : nth Prime 0 = 2 := nth_count prime_two +@[deprecated (since := "2024-09-21")] +alias zeroth_prime_eq_two := nth_prime_zero_eq_two + +@[simp] +theorem nth_prime_one_eq_three : nth Nat.Prime 1 = 3 := nth_count prime_three + +end Nat diff --git a/Mathlib/NumberTheory/PrimeCounting.lean b/Mathlib/NumberTheory/PrimeCounting.lean index a488afee078ab..fa1a732eddcb9 100644 --- a/Mathlib/NumberTheory/PrimeCounting.lean +++ b/Mathlib/NumberTheory/PrimeCounting.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Bolton Bailey. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Bolton Bailey, Ralf Stephan -/ -import Mathlib.Data.Nat.Nth +import Mathlib.Data.Nat.Prime.Nth import Mathlib.Data.Nat.Totient import Mathlib.NumberTheory.SmoothNumbers @@ -70,12 +70,9 @@ theorem monotone_primeCounting : Monotone primeCounting := theorem primeCounting'_nth_eq (n : ℕ) : π' (nth Prime n) = n := count_nth_of_infinite infinite_setOf_prime _ -@[simp] -theorem zeroth_prime_eq_two : nth Prime 0 = 2 := nth_count prime_two - /-- The `n`th prime is greater or equal to `n + 2`. -/ theorem add_two_le_nth_prime (n : ℕ) : n + 2 ≤ nth Prime n := - zeroth_prime_eq_two ▸ (nth_strictMono infinite_setOf_prime).add_le_nat n 0 + nth_prime_zero_eq_two ▸ (nth_strictMono infinite_setOf_prime).add_le_nat n 0 theorem surjective_primeCounting' : Function.Surjective π' := Nat.surjective_count_of_infinite_setOf infinite_setOf_prime @@ -99,6 +96,22 @@ theorem tensto_primeCounting : Tendsto π atTop atTop := theorem prime_nth_prime (n : ℕ) : Prime (nth Prime n) := nth_mem_of_infinite infinite_setOf_prime _ +@[simp] +lemma primeCounting'_eq_zero_iff {n : ℕ} : n.primeCounting' = 0 ↔ n ≤ 2 := by + rw [primeCounting', Nat.count_eq_zero ⟨_, Nat.prime_two⟩, Nat.nth_prime_zero_eq_two] + +@[simp] +lemma primeCounting_eq_zero_iff {n : ℕ} : n.primeCounting = 0 ↔ n ≤ 1 := by + simp [primeCounting] + +@[simp] +lemma primeCounting_zero : primeCounting 0 = 0 := + primeCounting_eq_zero_iff.mpr zero_le_one + +@[simp] +lemma primeCounting_one : primeCounting 1 = 0 := + primeCounting_eq_zero_iff.mpr le_rfl + /-- The cardinality of the finset `primesBelow n` equals the counting function `primeCounting'` at `n`. -/ theorem primesBelow_card_eq_primeCounting' (n : ℕ) : #n.primesBelow = primeCounting' n := by From ce45eadc5470ddcc2fc71a17fbec26a90bb8ef48 Mon Sep 17 00:00:00 2001 From: mathlib4-update-dependencies-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Mon, 28 Oct 2024 01:23:17 +0000 Subject: [PATCH 492/505] chore: update Mathlib dependencies 2024-10-28 (#18310) This PR updates the Mathlib dependencies. --- lake-manifest.json | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lake-manifest.json b/lake-manifest.json index cb37bf904e85f..2e679902d357f 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,7 +5,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "4d2cb85d87e0e728520fe0bf531cfdddb78fde7a", + "rev": "db2c5f8dd36bfab847ef42f5749cc4ba347be202", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", From 56c3b709d5ef80982afc96a34bf1c0a16a805cb5 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 28 Oct 2024 02:29:31 +0000 Subject: [PATCH 493/505] chore(Finite/Basic): act on a porting note (#18309) `infer_instance` no longer fails there. --- Mathlib/Data/Finite/Basic.lean | 5 +---- 1 file changed, 1 insertion(+), 4 deletions(-) diff --git a/Mathlib/Data/Finite/Basic.lean b/Mathlib/Data/Finite/Basic.lean index 59b373c1f2798..5ad53a591bc6e 100644 --- a/Mathlib/Data/Finite/Basic.lean +++ b/Mathlib/Data/Finite/Basic.lean @@ -60,10 +60,7 @@ instance Quotient.finite {α : Sort*} [Finite α] (s : Setoid α) : Finite (Quot instance Function.Embedding.finite {α β : Sort*} [Finite β] : Finite (α ↪ β) := by cases' isEmpty_or_nonempty (α ↪ β) with _ h - · -- Porting note: infer_instance fails because it applies `Finite.of_fintype` and produces a - -- "stuck at solving universe constraint" error. - apply Finite.of_subsingleton - + · infer_instance · refine h.elim fun f => ?_ haveI : Finite α := Finite.of_injective _ f.injective exact Finite.of_injective _ DFunLike.coe_injective From 132344bd6615720d726520986bd8dd66f0cbbaed Mon Sep 17 00:00:00 2001 From: mathlib4-update-dependencies-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Mon, 28 Oct 2024 03:13:07 +0000 Subject: [PATCH 494/505] chore: update Mathlib dependencies 2024-10-28 (#18314) This PR updates the Mathlib dependencies. --- lake-manifest.json | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lake-manifest.json b/lake-manifest.json index 2e679902d357f..89cb6852c1073 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,7 +5,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "db2c5f8dd36bfab847ef42f5749cc4ba347be202", + "rev": "85f6511d93f9e6a8188ec9985c82f08f65c26cae", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", From bafdce7177d4299610353f8fb55fec61de0fa294 Mon Sep 17 00:00:00 2001 From: Markus Himmel Date: Mon, 28 Oct 2024 04:19:21 +0000 Subject: [PATCH 495/505] feat(CategoryTheory): colimits after whiskering (#18282) --- .../Limits/FunctorCategory/Basic.lean | 46 +++++++++++++++++++ 1 file changed, 46 insertions(+) diff --git a/Mathlib/CategoryTheory/Limits/FunctorCategory/Basic.lean b/Mathlib/CategoryTheory/Limits/FunctorCategory/Basic.lean index 78e8c5a55bbdd..86b8cb71e4d54 100644 --- a/Mathlib/CategoryTheory/Limits/FunctorCategory/Basic.lean +++ b/Mathlib/CategoryTheory/Limits/FunctorCategory/Basic.lean @@ -236,6 +236,29 @@ theorem limit_obj_ext {H : J ⥤ K ⥤ C} [HasLimitsOfShape J C] {k : K} {W : C} ext j simpa using w j +/-- Taking a limit after whiskering by `G` is the same as using `G` and then taking a limit. -/ +def limitCompWhiskeringLeftIsoCompLimit (F : J ⥤ K ⥤ C) (G : D ⥤ K) [HasLimitsOfShape J C] : + limit (F ⋙ (whiskeringLeft _ _ _).obj G) ≅ G ⋙ limit F := + NatIso.ofComponents (fun j => + limitObjIsoLimitCompEvaluation (F ⋙ (whiskeringLeft _ _ _).obj G) j ≪≫ + HasLimit.isoOfNatIso (isoWhiskerLeft F (whiskeringLeftCompEvaluation G j)) ≪≫ + (limitObjIsoLimitCompEvaluation F (G.obj j)).symm) + +@[reassoc (attr := simp)] +theorem limitCompWhiskeringLeftIsoCompLimit_hom_whiskerLeft_π (F : J ⥤ K ⥤ C) (G : D ⥤ K) + [HasLimitsOfShape J C] (j : J) : + (limitCompWhiskeringLeftIsoCompLimit F G).hom ≫ whiskerLeft G (limit.π F j) = + limit.π (F ⋙ (whiskeringLeft _ _ _).obj G) j := by + ext d + simp [limitCompWhiskeringLeftIsoCompLimit] + +@[reassoc (attr := simp)] +theorem limitCompWhiskeringLeftIsoCompLimit_inv_π (F : J ⥤ K ⥤ C) (G : D ⥤ K) + [HasLimitsOfShape J C] (j : J) : + (limitCompWhiskeringLeftIsoCompLimit F G).inv ≫ limit.π (F ⋙ (whiskeringLeft _ _ _).obj G) j = + whiskerLeft G (limit.π F j) := by + simp [Iso.inv_comp_eq] + 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 _) @@ -302,6 +325,29 @@ theorem colimit_obj_ext {H : J ⥤ K ⥤ C} [HasColimitsOfShape J C] {k : K} {W ext j simpa using w j +/-- Taking a colimit after whiskering by `G` is the same as using `G` and then taking a colimit. -/ +def colimitCompWhiskeringLeftIsoCompColimit (F : J ⥤ K ⥤ C) (G : D ⥤ K) [HasColimitsOfShape J C] : + colimit (F ⋙ (whiskeringLeft _ _ _).obj G) ≅ G ⋙ colimit F := + NatIso.ofComponents (fun j => + colimitObjIsoColimitCompEvaluation (F ⋙ (whiskeringLeft _ _ _).obj G) j ≪≫ + HasColimit.isoOfNatIso (isoWhiskerLeft F (whiskeringLeftCompEvaluation G j)) ≪≫ + (colimitObjIsoColimitCompEvaluation F (G.obj j)).symm) + +@[reassoc (attr := simp)] +theorem ι_colimitCompWhiskeringLeftIsoCompColimit_hom (F : J ⥤ K ⥤ C) (G : D ⥤ K) + [HasColimitsOfShape J C] (j : J) : + colimit.ι (F ⋙ (whiskeringLeft _ _ _).obj G) j ≫ + (colimitCompWhiskeringLeftIsoCompColimit F G).hom = whiskerLeft G (colimit.ι F j) := by + ext d + simp [colimitCompWhiskeringLeftIsoCompColimit] + +@[reassoc (attr := simp)] +theorem whiskerLeft_ι_colimitCompWhiskeringLeftIsoCompColimit_inv (F : J ⥤ K ⥤ C) (G : D ⥤ K) + [HasColimitsOfShape J C] (j : J) : + whiskerLeft G (colimit.ι F j) ≫ (colimitCompWhiskeringLeftIsoCompColimit F G).inv = + colimit.ι (F ⋙ (whiskeringLeft _ _ _).obj G) j := by + simp [Iso.comp_inv_eq] + instance evaluationPreservesLimits [HasLimits C] (k : K) : PreservesLimits ((evaluation K C).obj k) where preservesLimitsOfShape {_} _𝒥 := inferInstance From 07221c86e111321ca6c27c5a2c435468d9610e66 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 28 Oct 2024 16:38:57 +1100 Subject: [PATCH 496/505] chore: cleanup skipAssignedInstances flags --- .../Algebra/Category/ModuleCat/ChangeOfRings.lean | 12 +++--------- Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean | 2 -- .../Homology/HomotopyCategory/DegreewiseSplit.lean | 3 +-- .../Homology/HomotopyCategory/Pretriangulated.lean | 2 -- Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean | 7 ++----- Mathlib/FieldTheory/KummerExtension.lean | 3 +-- Mathlib/MeasureTheory/Function/SimpleFunc.lean | 12 +++--------- Mathlib/MeasureTheory/Measure/Haar/OfBasis.lean | 2 -- Mathlib/Order/Closure.lean | 2 +- Mathlib/Order/Filter/Basic.lean | 2 +- Mathlib/Probability/Process/Stopping.lean | 3 +-- .../Topology/Category/TopCat/Limits/Products.lean | 4 +--- 12 files changed, 14 insertions(+), 40 deletions(-) diff --git a/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean b/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean index 9d1934b7c1d11..b74a03abb8658 100644 --- a/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean +++ b/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean @@ -808,21 +808,15 @@ def extendRestrictScalarsAdj {R : Type u₁} {S : Type u₂} [CommRing R] [CommR rw [ModuleCat.coe_comp, Function.comp_apply, restrictScalars.map_apply] rfl homEquiv_counit := fun {X Y g} ↦ LinearMap.ext fun x => by - -- Porting note: once again reminding Lean of the instances - letI m1 : Module R S := Module.compHom S f - letI m2 : Module R Y := Module.compHom Y f induction x using TensorProduct.induction_on with | zero => rw [map_zero, map_zero] | tmul => rw [ExtendRestrictScalarsAdj.homEquiv_symm_apply] - -- This used to be `rw`, but we need `erw` after leanprover/lean4#2644 - erw [ModuleCat.coe_comp] - rw [Function.comp_apply, ExtendRestrictScalarsAdj.counit_app] + dsimp + rw [ModuleCat.coe_comp, Function.comp_apply] -- This used to be `rw`, but we need `erw` after leanprover/lean4#2644 erw [ExtendRestrictScalarsAdj.Counit.map_apply] - set_option tactic.skipAssignedInstances false in dsimp - rw [TensorProduct.lift.tmul] - rfl + dsimp | add => rw [map_add, map_add]; congr 1 } instance {R : Type u₁} {S : Type u₂} [CommRing R] [CommRing S] (f : R →+* S) : diff --git a/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean b/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean index 840cc5164aa40..1f5104d65643a 100644 --- a/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean +++ b/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean @@ -982,8 +982,6 @@ theorem mem_prod_list_ofFn {a : α} {s : Fin n → Finset α} : @[to_additive] theorem mem_pow {a : α} {n : ℕ} : a ∈ s ^ n ↔ ∃ f : Fin n → s, (List.ofFn fun i => ↑(f i)).prod = a := by - -- Also compiles without the option, but much slower. - set_option tactic.skipAssignedInstances false in simp [← mem_coe, coe_pow, Set.mem_pow] @[to_additive (attr := simp) nsmul_empty] diff --git a/Mathlib/Algebra/Homology/HomotopyCategory/DegreewiseSplit.lean b/Mathlib/Algebra/Homology/HomotopyCategory/DegreewiseSplit.lean index a8ad46dc1d626..a4260ebc14480 100644 --- a/Mathlib/Algebra/Homology/HomotopyCategory/DegreewiseSplit.lean +++ b/Mathlib/Algebra/Homology/HomotopyCategory/DegreewiseSplit.lean @@ -70,7 +70,6 @@ noncomputable abbrev trianglehOfDegreewiseSplit : variable [HasBinaryBiproducts C] -set_option tactic.skipAssignedInstances false in /-- The canonical isomorphism `(mappingCone (homOfDegreewiseSplit S σ)).X p ≅ S.X₂.X q` when `p + 1 = q`. -/ noncomputable def mappingConeHomOfDegreewiseSplitXIso (p q : ℤ) (hpq : p + 1 = q) : @@ -104,6 +103,7 @@ noncomputable def mappingConeHomOfDegreewiseSplitXIso (p q : ℤ) (hpq : p + 1 = mappingCone.inl_v_snd_v_assoc, mappingCone.inr_f_snd_v_assoc, zero_sub, sub_neg_eq_add, ← h] abel +-- Works without this `set_option`, but takes about 20% longer. set_option backward.isDefEq.lazyWhnfCore false in -- See https://github.com/leanprover-community/mathlib4/issues/12534 /-- The canonical isomorphism `mappingCone (homOfDegreewiseSplit S σ) ≅ S.X₂⟦(1 : ℤ)⟧`. -/ @[simps!] @@ -114,7 +114,6 @@ noncomputable def mappingConeHomOfDegreewiseSplitIso : have r_f := (σ (p + 1 + 1)).r_f have s_g := (σ (p + 1)).s_g dsimp at r_f s_g - set_option tactic.skipAssignedInstances false in simp [mappingConeHomOfDegreewiseSplitXIso, mappingCone.ext_from_iff _ _ _ rfl, mappingCone.inl_v_d_assoc _ (p + 1) _ (p + 1 + 1) (by linarith) (by linarith), cocycleOfDegreewiseSplit, r_f] diff --git a/Mathlib/Algebra/Homology/HomotopyCategory/Pretriangulated.lean b/Mathlib/Algebra/Homology/HomotopyCategory/Pretriangulated.lean index d5e210fe2d0c1..78e5edc499ba5 100644 --- a/Mathlib/Algebra/Homology/HomotopyCategory/Pretriangulated.lean +++ b/Mathlib/Algebra/Homology/HomotopyCategory/Pretriangulated.lean @@ -323,13 +323,11 @@ noncomputable def shiftTriangleIso (n : ℤ) : Triangle.mk_obj₂, Triangle.mk_mor₁, Preadditive.smul_iso_hom, Iso.refl_hom, Linear.comp_smul, comp_id, smul_smul, Int.units_coe_mul_self, one_smul, id_comp] · ext p - set_option tactic.skipAssignedInstances false in dsimp simp only [Units.smul_def, shiftIso, Int.reduceNeg, Linear.smul_comp, id_comp, ext_to_iff _ _ (p + 1) rfl, shiftFunctor_obj_X', assoc, lift_f_fst_v, Cocycle.coe_smul, Cocycle.shift_coe, Cochain.smul_v, Cochain.shift_v', Linear.comp_smul, inr_f_fst_v, smul_zero, lift_f_snd_v, inr_f_snd_v, and_true] - rw [smul_zero] · ext p dsimp simp only [triangle, Triangle.mk_mor₃, Cocycle.homOf_f, Cocycle.rightShift_coe, diff --git a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean index 2ad92e90d3efa..631c55eebc791 100644 --- a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean +++ b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean @@ -272,11 +272,8 @@ open Function RingHom theorem comap_inducing_of_surjective (hf : Surjective f) : Inducing (comap f) where induced := by - set_option tactic.skipAssignedInstances false in - simp_rw [TopologicalSpace.ext_iff, ← isClosed_compl_iff, - ← @isClosed_compl_iff (PrimeSpectrum S) - ((TopologicalSpace.induced (comap f) zariskiTopology)), isClosed_induced_iff, - isClosed_iff_zeroLocus] + simp only [TopologicalSpace.ext_iff, ← isClosed_compl_iff, isClosed_iff_zeroLocus, + isClosed_induced_iff] refine fun s => ⟨fun ⟨F, hF⟩ => ⟨zeroLocus (f ⁻¹' F), ⟨f ⁻¹' F, rfl⟩, by diff --git a/Mathlib/FieldTheory/KummerExtension.lean b/Mathlib/FieldTheory/KummerExtension.lean index bb7dbd753c5b8..095253e93171d 100644 --- a/Mathlib/FieldTheory/KummerExtension.lean +++ b/Mathlib/FieldTheory/KummerExtension.lean @@ -369,7 +369,6 @@ lemma autAdjoinRootXPowSubCEquiv_root (η) : autAdjoinRootXPowSubCEquiv hζ hn H η (root _) = ((η : Kˣ) : K) • root _ := autAdjoinRootXPowSubC_root hn a η -set_option tactic.skipAssignedInstances false in lemma autAdjoinRootXPowSubCEquiv_symm_smul (σ) : ((autAdjoinRootXPowSubCEquiv hζ hn H).symm σ : Kˣ) • (root _ : K[n√a]) = σ (root _) := by have := Fact.mk H @@ -380,7 +379,7 @@ lemma autAdjoinRootXPowSubCEquiv_symm_smul (σ) : Units.val_ofPowEqOne, ite_mul, one_mul, ne_eq] simp_rw [← root_X_pow_sub_C_eq_zero_iff H] split_ifs with h - · rw [h, mul_zero, map_zero] + · rw [h, map_zero] · rw [div_mul_cancel₀ _ h] end AdjoinRoot diff --git a/Mathlib/MeasureTheory/Function/SimpleFunc.lean b/Mathlib/MeasureTheory/Function/SimpleFunc.lean index 0a08303ca10a2..703145555d7ac 100644 --- a/Mathlib/MeasureTheory/Function/SimpleFunc.lean +++ b/Mathlib/MeasureTheory/Function/SimpleFunc.lean @@ -197,21 +197,15 @@ theorem piecewise_apply {s : Set α} (hs : MeasurableSet s) (f g : α →ₛ β) @[simp] theorem piecewise_compl {s : Set α} (hs : MeasurableSet sᶜ) (f g : α →ₛ β) : piecewise sᶜ hs f g = piecewise s hs.of_compl g f := - coe_injective <| by - set_option tactic.skipAssignedInstances false in - simp [hs]; convert Set.piecewise_compl s f g + coe_injective <| by simp [hs] @[simp] theorem piecewise_univ (f g : α →ₛ β) : piecewise univ MeasurableSet.univ f g = f := - coe_injective <| by - set_option tactic.skipAssignedInstances false in - simp; convert Set.piecewise_univ f g + coe_injective <| by simp @[simp] theorem piecewise_empty (f g : α →ₛ β) : piecewise ∅ MeasurableSet.empty f g = g := - coe_injective <| by - set_option tactic.skipAssignedInstances false in - simp; convert Set.piecewise_empty f g + coe_injective <| by simp @[simp] theorem piecewise_same (f : α →ₛ β) {s : Set α} (hs : MeasurableSet s) : diff --git a/Mathlib/MeasureTheory/Measure/Haar/OfBasis.lean b/Mathlib/MeasureTheory/Measure/Haar/OfBasis.lean index 853e3faeab186..7f76b5cbe8048 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/OfBasis.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/OfBasis.lean @@ -213,8 +213,6 @@ theorem Basis.parallelepiped_map (b : Basis ι ℝ E) (e : E ≃ₗ[ℝ] F) : LinearMap.isOpenMap_of_finiteDimensional _ e.surjective) := PositiveCompacts.ext (image_parallelepiped e.toLinearMap _).symm --- removing this option makes elaboration approximately 1 second slower -set_option tactic.skipAssignedInstances false in theorem Basis.prod_parallelepiped (v : Basis ι ℝ E) (w : Basis ι' ℝ F) : (v.prod w).parallelepiped = v.parallelepiped.prod w.parallelepiped := by ext x diff --git a/Mathlib/Order/Closure.lean b/Mathlib/Order/Closure.lean index ab229a68e246f..0759232ea5236 100644 --- a/Mathlib/Order/Closure.lean +++ b/Mathlib/Order/Closure.lean @@ -254,7 +254,7 @@ variable [CompleteLattice α] (c : ClosureOperator α) @[simps!] def ofCompletePred (p : α → Prop) (hsinf : ∀ s, (∀ a ∈ s, p a) → p (sInf s)) : ClosureOperator α := ofPred (fun a ↦ ⨅ b : {b // a ≤ b ∧ p b}, b) p - (fun a ↦ by set_option tactic.skipAssignedInstances false in simp [forall_swap]) + (fun a ↦ by simp [forall_swap]) (fun _ ↦ hsinf _ <| forall_mem_range.2 fun b ↦ b.2.2) (fun _ b hab hb ↦ iInf_le_of_le ⟨b, hab, hb⟩ le_rfl) diff --git a/Mathlib/Order/Filter/Basic.lean b/Mathlib/Order/Filter/Basic.lean index 47df88df7daec..3587c2a0fa469 100644 --- a/Mathlib/Order/Filter/Basic.lean +++ b/Mathlib/Order/Filter/Basic.lean @@ -991,7 +991,7 @@ theorem frequently_imp_distrib_left {f : Filter α} [NeBot f] {p : Prop} {q : α theorem frequently_imp_distrib_right {f : Filter α} [NeBot f] {p : α → Prop} {q : Prop} : (∃ᶠ x in f, p x → q) ↔ (∀ᶠ x in f, p x) → q := by - set_option tactic.skipAssignedInstances false in simp [frequently_imp_distrib] + simp only [frequently_imp_distrib, frequently_const] theorem eventually_imp_distrib_right {f : Filter α} {p : α → Prop} {q : Prop} : (∀ᶠ x in f, p x → q) ↔ (∃ᶠ x in f, p x) → q := by diff --git a/Mathlib/Probability/Process/Stopping.lean b/Mathlib/Probability/Process/Stopping.lean index c9b88bfbb13b3..40bf2069437c8 100644 --- a/Mathlib/Probability/Process/Stopping.lean +++ b/Mathlib/Probability/Process/Stopping.lean @@ -387,8 +387,7 @@ theorem measurableSet_inter_eq_iff (hτ : IsStoppingTime f τ) (s : Set Ω) (i : by_cases hij : i ≤ j · simp only [hij, Set.setOf_true, Set.inter_univ] exact f.mono hij _ h - · set_option tactic.skipAssignedInstances false in simp [hij] - convert @MeasurableSet.empty _ (Filtration.seq f j) + · simp [hij] theorem measurableSpace_le_of_le_const (hτ : IsStoppingTime f τ) {i : ι} (hτ_le : ∀ ω, τ ω ≤ i) : hτ.measurableSpace ≤ f i := diff --git a/Mathlib/Topology/Category/TopCat/Limits/Products.lean b/Mathlib/Topology/Category/TopCat/Limits/Products.lean index 132e2144e2090..debcb36fa5510 100644 --- a/Mathlib/Topology/Category/TopCat/Limits/Products.lean +++ b/Mathlib/Topology/Category/TopCat/Limits/Products.lean @@ -47,9 +47,7 @@ def piFanIsLimit {ι : Type v} (α : ι → TopCat.{max v u}) : IsLimit (piFan intro S m h apply ContinuousMap.ext; intro x funext i - set_option tactic.skipAssignedInstances false in - dsimp - rw [ContinuousMap.coe_mk, ← h ⟨i⟩] + simp [ContinuousMap.coe_mk, ← h ⟨i⟩] rfl fac _ _ := rfl From 3458918ae1c7d3898f1493a7df2291f4a9f74b34 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 28 Oct 2024 16:39:36 +1100 Subject: [PATCH 497/505] Revert "chore: cleanup skipAssignedInstances flags" This reverts commit 07221c86e111321ca6c27c5a2c435468d9610e66. --- .../Algebra/Category/ModuleCat/ChangeOfRings.lean | 12 +++++++++--- Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean | 2 ++ .../Homology/HomotopyCategory/DegreewiseSplit.lean | 3 ++- .../Homology/HomotopyCategory/Pretriangulated.lean | 2 ++ Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean | 7 +++++-- Mathlib/FieldTheory/KummerExtension.lean | 3 ++- Mathlib/MeasureTheory/Function/SimpleFunc.lean | 12 +++++++++--- Mathlib/MeasureTheory/Measure/Haar/OfBasis.lean | 2 ++ Mathlib/Order/Closure.lean | 2 +- Mathlib/Order/Filter/Basic.lean | 2 +- Mathlib/Probability/Process/Stopping.lean | 3 ++- .../Topology/Category/TopCat/Limits/Products.lean | 4 +++- 12 files changed, 40 insertions(+), 14 deletions(-) diff --git a/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean b/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean index b74a03abb8658..9d1934b7c1d11 100644 --- a/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean +++ b/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean @@ -808,15 +808,21 @@ def extendRestrictScalarsAdj {R : Type u₁} {S : Type u₂} [CommRing R] [CommR rw [ModuleCat.coe_comp, Function.comp_apply, restrictScalars.map_apply] rfl homEquiv_counit := fun {X Y g} ↦ LinearMap.ext fun x => by + -- Porting note: once again reminding Lean of the instances + letI m1 : Module R S := Module.compHom S f + letI m2 : Module R Y := Module.compHom Y f induction x using TensorProduct.induction_on with | zero => rw [map_zero, map_zero] | tmul => rw [ExtendRestrictScalarsAdj.homEquiv_symm_apply] - dsimp - rw [ModuleCat.coe_comp, Function.comp_apply] + -- This used to be `rw`, but we need `erw` after leanprover/lean4#2644 + erw [ModuleCat.coe_comp] + rw [Function.comp_apply, ExtendRestrictScalarsAdj.counit_app] -- This used to be `rw`, but we need `erw` after leanprover/lean4#2644 erw [ExtendRestrictScalarsAdj.Counit.map_apply] - dsimp + set_option tactic.skipAssignedInstances false in dsimp + rw [TensorProduct.lift.tmul] + rfl | add => rw [map_add, map_add]; congr 1 } instance {R : Type u₁} {S : Type u₂} [CommRing R] [CommRing S] (f : R →+* S) : diff --git a/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean b/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean index 1f5104d65643a..840cc5164aa40 100644 --- a/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean +++ b/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean @@ -982,6 +982,8 @@ theorem mem_prod_list_ofFn {a : α} {s : Fin n → Finset α} : @[to_additive] theorem mem_pow {a : α} {n : ℕ} : a ∈ s ^ n ↔ ∃ f : Fin n → s, (List.ofFn fun i => ↑(f i)).prod = a := by + -- Also compiles without the option, but much slower. + set_option tactic.skipAssignedInstances false in simp [← mem_coe, coe_pow, Set.mem_pow] @[to_additive (attr := simp) nsmul_empty] diff --git a/Mathlib/Algebra/Homology/HomotopyCategory/DegreewiseSplit.lean b/Mathlib/Algebra/Homology/HomotopyCategory/DegreewiseSplit.lean index a4260ebc14480..a8ad46dc1d626 100644 --- a/Mathlib/Algebra/Homology/HomotopyCategory/DegreewiseSplit.lean +++ b/Mathlib/Algebra/Homology/HomotopyCategory/DegreewiseSplit.lean @@ -70,6 +70,7 @@ noncomputable abbrev trianglehOfDegreewiseSplit : variable [HasBinaryBiproducts C] +set_option tactic.skipAssignedInstances false in /-- The canonical isomorphism `(mappingCone (homOfDegreewiseSplit S σ)).X p ≅ S.X₂.X q` when `p + 1 = q`. -/ noncomputable def mappingConeHomOfDegreewiseSplitXIso (p q : ℤ) (hpq : p + 1 = q) : @@ -103,7 +104,6 @@ noncomputable def mappingConeHomOfDegreewiseSplitXIso (p q : ℤ) (hpq : p + 1 = mappingCone.inl_v_snd_v_assoc, mappingCone.inr_f_snd_v_assoc, zero_sub, sub_neg_eq_add, ← h] abel --- Works without this `set_option`, but takes about 20% longer. set_option backward.isDefEq.lazyWhnfCore false in -- See https://github.com/leanprover-community/mathlib4/issues/12534 /-- The canonical isomorphism `mappingCone (homOfDegreewiseSplit S σ) ≅ S.X₂⟦(1 : ℤ)⟧`. -/ @[simps!] @@ -114,6 +114,7 @@ noncomputable def mappingConeHomOfDegreewiseSplitIso : have r_f := (σ (p + 1 + 1)).r_f have s_g := (σ (p + 1)).s_g dsimp at r_f s_g + set_option tactic.skipAssignedInstances false in simp [mappingConeHomOfDegreewiseSplitXIso, mappingCone.ext_from_iff _ _ _ rfl, mappingCone.inl_v_d_assoc _ (p + 1) _ (p + 1 + 1) (by linarith) (by linarith), cocycleOfDegreewiseSplit, r_f] diff --git a/Mathlib/Algebra/Homology/HomotopyCategory/Pretriangulated.lean b/Mathlib/Algebra/Homology/HomotopyCategory/Pretriangulated.lean index 78e5edc499ba5..d5e210fe2d0c1 100644 --- a/Mathlib/Algebra/Homology/HomotopyCategory/Pretriangulated.lean +++ b/Mathlib/Algebra/Homology/HomotopyCategory/Pretriangulated.lean @@ -323,11 +323,13 @@ noncomputable def shiftTriangleIso (n : ℤ) : Triangle.mk_obj₂, Triangle.mk_mor₁, Preadditive.smul_iso_hom, Iso.refl_hom, Linear.comp_smul, comp_id, smul_smul, Int.units_coe_mul_self, one_smul, id_comp] · ext p + set_option tactic.skipAssignedInstances false in dsimp simp only [Units.smul_def, shiftIso, Int.reduceNeg, Linear.smul_comp, id_comp, ext_to_iff _ _ (p + 1) rfl, shiftFunctor_obj_X', assoc, lift_f_fst_v, Cocycle.coe_smul, Cocycle.shift_coe, Cochain.smul_v, Cochain.shift_v', Linear.comp_smul, inr_f_fst_v, smul_zero, lift_f_snd_v, inr_f_snd_v, and_true] + rw [smul_zero] · ext p dsimp simp only [triangle, Triangle.mk_mor₃, Cocycle.homOf_f, Cocycle.rightShift_coe, diff --git a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean index 631c55eebc791..2ad92e90d3efa 100644 --- a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean +++ b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean @@ -272,8 +272,11 @@ open Function RingHom theorem comap_inducing_of_surjective (hf : Surjective f) : Inducing (comap f) where induced := by - simp only [TopologicalSpace.ext_iff, ← isClosed_compl_iff, isClosed_iff_zeroLocus, - isClosed_induced_iff] + set_option tactic.skipAssignedInstances false in + simp_rw [TopologicalSpace.ext_iff, ← isClosed_compl_iff, + ← @isClosed_compl_iff (PrimeSpectrum S) + ((TopologicalSpace.induced (comap f) zariskiTopology)), isClosed_induced_iff, + isClosed_iff_zeroLocus] refine fun s => ⟨fun ⟨F, hF⟩ => ⟨zeroLocus (f ⁻¹' F), ⟨f ⁻¹' F, rfl⟩, by diff --git a/Mathlib/FieldTheory/KummerExtension.lean b/Mathlib/FieldTheory/KummerExtension.lean index 095253e93171d..bb7dbd753c5b8 100644 --- a/Mathlib/FieldTheory/KummerExtension.lean +++ b/Mathlib/FieldTheory/KummerExtension.lean @@ -369,6 +369,7 @@ lemma autAdjoinRootXPowSubCEquiv_root (η) : autAdjoinRootXPowSubCEquiv hζ hn H η (root _) = ((η : Kˣ) : K) • root _ := autAdjoinRootXPowSubC_root hn a η +set_option tactic.skipAssignedInstances false in lemma autAdjoinRootXPowSubCEquiv_symm_smul (σ) : ((autAdjoinRootXPowSubCEquiv hζ hn H).symm σ : Kˣ) • (root _ : K[n√a]) = σ (root _) := by have := Fact.mk H @@ -379,7 +380,7 @@ lemma autAdjoinRootXPowSubCEquiv_symm_smul (σ) : Units.val_ofPowEqOne, ite_mul, one_mul, ne_eq] simp_rw [← root_X_pow_sub_C_eq_zero_iff H] split_ifs with h - · rw [h, map_zero] + · rw [h, mul_zero, map_zero] · rw [div_mul_cancel₀ _ h] end AdjoinRoot diff --git a/Mathlib/MeasureTheory/Function/SimpleFunc.lean b/Mathlib/MeasureTheory/Function/SimpleFunc.lean index 703145555d7ac..0a08303ca10a2 100644 --- a/Mathlib/MeasureTheory/Function/SimpleFunc.lean +++ b/Mathlib/MeasureTheory/Function/SimpleFunc.lean @@ -197,15 +197,21 @@ theorem piecewise_apply {s : Set α} (hs : MeasurableSet s) (f g : α →ₛ β) @[simp] theorem piecewise_compl {s : Set α} (hs : MeasurableSet sᶜ) (f g : α →ₛ β) : piecewise sᶜ hs f g = piecewise s hs.of_compl g f := - coe_injective <| by simp [hs] + coe_injective <| by + set_option tactic.skipAssignedInstances false in + simp [hs]; convert Set.piecewise_compl s f g @[simp] theorem piecewise_univ (f g : α →ₛ β) : piecewise univ MeasurableSet.univ f g = f := - coe_injective <| by simp + coe_injective <| by + set_option tactic.skipAssignedInstances false in + simp; convert Set.piecewise_univ f g @[simp] theorem piecewise_empty (f g : α →ₛ β) : piecewise ∅ MeasurableSet.empty f g = g := - coe_injective <| by simp + coe_injective <| by + set_option tactic.skipAssignedInstances false in + simp; convert Set.piecewise_empty f g @[simp] theorem piecewise_same (f : α →ₛ β) {s : Set α} (hs : MeasurableSet s) : diff --git a/Mathlib/MeasureTheory/Measure/Haar/OfBasis.lean b/Mathlib/MeasureTheory/Measure/Haar/OfBasis.lean index 7f76b5cbe8048..853e3faeab186 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/OfBasis.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/OfBasis.lean @@ -213,6 +213,8 @@ theorem Basis.parallelepiped_map (b : Basis ι ℝ E) (e : E ≃ₗ[ℝ] F) : LinearMap.isOpenMap_of_finiteDimensional _ e.surjective) := PositiveCompacts.ext (image_parallelepiped e.toLinearMap _).symm +-- removing this option makes elaboration approximately 1 second slower +set_option tactic.skipAssignedInstances false in theorem Basis.prod_parallelepiped (v : Basis ι ℝ E) (w : Basis ι' ℝ F) : (v.prod w).parallelepiped = v.parallelepiped.prod w.parallelepiped := by ext x diff --git a/Mathlib/Order/Closure.lean b/Mathlib/Order/Closure.lean index 0759232ea5236..ab229a68e246f 100644 --- a/Mathlib/Order/Closure.lean +++ b/Mathlib/Order/Closure.lean @@ -254,7 +254,7 @@ variable [CompleteLattice α] (c : ClosureOperator α) @[simps!] def ofCompletePred (p : α → Prop) (hsinf : ∀ s, (∀ a ∈ s, p a) → p (sInf s)) : ClosureOperator α := ofPred (fun a ↦ ⨅ b : {b // a ≤ b ∧ p b}, b) p - (fun a ↦ by simp [forall_swap]) + (fun a ↦ by set_option tactic.skipAssignedInstances false in simp [forall_swap]) (fun _ ↦ hsinf _ <| forall_mem_range.2 fun b ↦ b.2.2) (fun _ b hab hb ↦ iInf_le_of_le ⟨b, hab, hb⟩ le_rfl) diff --git a/Mathlib/Order/Filter/Basic.lean b/Mathlib/Order/Filter/Basic.lean index 3587c2a0fa469..47df88df7daec 100644 --- a/Mathlib/Order/Filter/Basic.lean +++ b/Mathlib/Order/Filter/Basic.lean @@ -991,7 +991,7 @@ theorem frequently_imp_distrib_left {f : Filter α} [NeBot f] {p : Prop} {q : α theorem frequently_imp_distrib_right {f : Filter α} [NeBot f] {p : α → Prop} {q : Prop} : (∃ᶠ x in f, p x → q) ↔ (∀ᶠ x in f, p x) → q := by - simp only [frequently_imp_distrib, frequently_const] + set_option tactic.skipAssignedInstances false in simp [frequently_imp_distrib] theorem eventually_imp_distrib_right {f : Filter α} {p : α → Prop} {q : Prop} : (∀ᶠ x in f, p x → q) ↔ (∃ᶠ x in f, p x) → q := by diff --git a/Mathlib/Probability/Process/Stopping.lean b/Mathlib/Probability/Process/Stopping.lean index 40bf2069437c8..c9b88bfbb13b3 100644 --- a/Mathlib/Probability/Process/Stopping.lean +++ b/Mathlib/Probability/Process/Stopping.lean @@ -387,7 +387,8 @@ theorem measurableSet_inter_eq_iff (hτ : IsStoppingTime f τ) (s : Set Ω) (i : by_cases hij : i ≤ j · simp only [hij, Set.setOf_true, Set.inter_univ] exact f.mono hij _ h - · simp [hij] + · set_option tactic.skipAssignedInstances false in simp [hij] + convert @MeasurableSet.empty _ (Filtration.seq f j) theorem measurableSpace_le_of_le_const (hτ : IsStoppingTime f τ) {i : ι} (hτ_le : ∀ ω, τ ω ≤ i) : hτ.measurableSpace ≤ f i := diff --git a/Mathlib/Topology/Category/TopCat/Limits/Products.lean b/Mathlib/Topology/Category/TopCat/Limits/Products.lean index debcb36fa5510..132e2144e2090 100644 --- a/Mathlib/Topology/Category/TopCat/Limits/Products.lean +++ b/Mathlib/Topology/Category/TopCat/Limits/Products.lean @@ -47,7 +47,9 @@ def piFanIsLimit {ι : Type v} (α : ι → TopCat.{max v u}) : IsLimit (piFan intro S m h apply ContinuousMap.ext; intro x funext i - simp [ContinuousMap.coe_mk, ← h ⟨i⟩] + set_option tactic.skipAssignedInstances false in + dsimp + rw [ContinuousMap.coe_mk, ← h ⟨i⟩] rfl fac _ _ := rfl From 174e4bd31d28b82604fc68a45c04fbc15140c394 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 28 Oct 2024 06:12:34 +0000 Subject: [PATCH 498/505] =?UTF-8?q?feat:=20The=20Erd=C5=91s=E2=80=93Ko?= =?UTF-8?q?=E2=80=93Rado=20theorem=20(#15705)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Bhavik Mehta --- .../SetFamily/KruskalKatona.lean | 100 ++++++++++++++++++ 1 file changed, 100 insertions(+) diff --git a/Mathlib/Combinatorics/SetFamily/KruskalKatona.lean b/Mathlib/Combinatorics/SetFamily/KruskalKatona.lean index f8957f5d2c015..991e4071cf4bf 100644 --- a/Mathlib/Combinatorics/SetFamily/KruskalKatona.lean +++ b/Mathlib/Combinatorics/SetFamily/KruskalKatona.lean @@ -5,6 +5,8 @@ Authors: Bhavik Mehta, Yaël Dillies -/ import Mathlib.Combinatorics.Colex import Mathlib.Combinatorics.SetFamily.Compression.UV +import Mathlib.Combinatorics.SetFamily.Intersecting +import Mathlib.Data.Finset.Fin /-! # Kruskal-Katona theorem @@ -290,5 +292,103 @@ theorem iterated_kk (h₁ : (𝒜 : Set (Finset (Fin n))).Sized r) (h₂ : #𝒞 · refine ih h₁.shadow (kruskal_katona h₁ h₂ h₃) ?_ convert h₃.shadow +/-- The **Lovasz formulation of the Kruskal-Katona theorem**. + +If `|𝒜| ≥ k choose r`, (and everything in `𝒜` has size `r`) then the initial segment we compare to +is just all the subsets of `{0, ..., k - 1}` of size `r`. The `i`-th iterated shadow of this is all +the subsets of `{0, ..., k - 1}` of size `r - i`, so the `i`-th iterated shadow of `𝒜` has at least +`k.choose (r - i)` elements. -/ +theorem kruskal_katona_lovasz_form (hir : i ≤ r) (hrk : r ≤ k) (hkn : k ≤ n) + (h₁ : (𝒜 : Set (Finset (Fin n))).Sized r) (h₂ : k.choose r ≤ #𝒜) : + k.choose (r - i) ≤ #(∂^[i] 𝒜) := by + set range'k : Finset (Fin n) := + attachFin (range k) fun m ↦ by rw [mem_range]; apply forall_lt_iff_le.2 hkn + set 𝒞 : Finset (Finset (Fin n)) := powersetCard r range'k + have : (𝒞 : Set (Finset (Fin n))).Sized r := Set.sized_powersetCard _ _ + calc + k.choose (r - i) + _ = #(powersetCard (r - i) range'k) := by rw [card_powersetCard, card_attachFin, card_range] + _ = #(∂^[i] 𝒞) := by + congr! + ext B + rw [mem_powersetCard, mem_shadow_iterate_iff_exists_sdiff] + constructor + · rintro ⟨hBk, hB⟩ + have := exists_subsuperset_card_eq hBk (Nat.le_add_left _ i) <| by + rwa [hB, card_attachFin, card_range, ← Nat.add_sub_assoc hir, Nat.add_sub_cancel_left] + obtain ⟨C, BsubC, hCrange, hcard⟩ := this + rw [hB, ← Nat.add_sub_assoc hir, Nat.add_sub_cancel_left] at hcard + refine ⟨C, mem_powersetCard.2 ⟨hCrange, hcard⟩, BsubC, ?_⟩ + rw [card_sdiff BsubC, hcard, hB, Nat.sub_sub_self hir] + · rintro ⟨A, Ah, hBA, card_sdiff_i⟩ + rw [mem_powersetCard] at Ah + refine ⟨hBA.trans Ah.1, eq_tsub_of_add_eq ?_⟩ + rw [← Ah.2, ← card_sdiff_i, add_comm, card_sdiff_add_card_eq_card hBA] + _ ≤ #(∂ ^[i] 𝒜) := by + refine iterated_kk h₁ ?_ ⟨‹_›, ?_⟩ + · rwa [card_powersetCard, card_attachFin, card_range] + simp_rw [𝒞, mem_powersetCard] + rintro A B hA ⟨HB₁, HB₂⟩ + refine ⟨fun t ht ↦ ?_, ‹_›⟩ + rw [mem_attachFin, mem_range] + have : toColex (image Fin.val B) < toColex (image Fin.val A) := by + rwa [toColex_image_lt_toColex_image Fin.val_strictMono] + apply Colex.forall_lt_mono this.le _ t (mem_image.2 ⟨t, ht, rfl⟩) + simp_rw [mem_image] + rintro _ ⟨a, ha, hab⟩ + simpa [range'k, hab] using hA.1 ha + end KK + +/-- The **Erdős–Ko–Rado theorem**. + +The maximum size of an intersecting family in `α` where all sets have size `r` is bounded by +`(card α - 1).choose (r - 1)`. This bound is sharp. -/ +theorem erdos_ko_rado {𝒜 : Finset (Finset (Fin n))} {r : ℕ} + (h𝒜 : (𝒜 : Set (Finset (Fin n))).Intersecting) (h₂ : (𝒜 : Set (Finset (Fin n))).Sized r) + (h₃ : r ≤ n / 2) : + #𝒜 ≤ (n - 1).choose (r - 1) := by + -- Take care of the r=0 case first: it's not very interesting. + cases' Nat.eq_zero_or_pos r with b h1r + · convert Nat.zero_le _ + rw [Finset.card_eq_zero, eq_empty_iff_forall_not_mem] + refine fun A HA ↦ h𝒜 HA HA ?_ + rw [disjoint_self_iff_empty, ← Finset.card_eq_zero, ← b] + exact h₂ HA + refine le_of_not_lt fun size ↦ ?_ + -- Consider 𝒜ᶜˢ = {sᶜ | s ∈ 𝒜} + -- Its iterated shadow (∂^[n-2k] 𝒜ᶜˢ) is disjoint from 𝒜 by intersecting-ness + have : Disjoint 𝒜 (∂^[n - 2 * r] 𝒜ᶜˢ) := disjoint_right.2 fun A hAbar hA ↦ by + simp [mem_shadow_iterate_iff_exists_sdiff, mem_compls] at hAbar + obtain ⟨C, hC, hAC, _⟩ := hAbar + exact h𝒜 hA hC (disjoint_of_subset_left hAC disjoint_compl_right) + have : r ≤ n := h₃.trans (Nat.div_le_self n 2) + have : 1 ≤ n := ‹1 ≤ r›.trans ‹r ≤ n› + -- We know the size of 𝒜ᶜˢ since it's the same size as 𝒜 + have z : (n - 1).choose (n - r) < #𝒜ᶜˢ := by + rwa [card_compls, choose_symm_of_eq_add (tsub_add_tsub_cancel ‹r ≤ n› ‹1 ≤ r›).symm] + -- and everything in 𝒜ᶜˢ has size n-r. + have h𝒜bar : (𝒜ᶜˢ : Set (Finset (Fin n))).Sized (n - r) := by simpa using h₂.compls + have : n - 2 * r ≤ n - r := by + rw [tsub_le_tsub_iff_left ‹r ≤ n›] + exact Nat.le_mul_of_pos_left _ zero_lt_two + -- We can use the Lovasz form of Kruskal-Katona to get |∂^[n-2k] 𝒜ᶜˢ| ≥ (n-1) choose r + have kk := kruskal_katona_lovasz_form ‹n - 2 * r ≤ n - r› ((tsub_le_tsub_iff_left ‹1 ≤ n›).2 h1r) + tsub_le_self h𝒜bar z.le + have : n - r - (n - 2 * r) = r := by omega + rw [this] at kk + -- But this gives a contradiction: `n choose r < |𝒜| + |∂^[n-2k] 𝒜ᶜˢ|` + have : n.choose r < #(𝒜 ∪ ∂^[n - 2 * r] 𝒜ᶜˢ) := by + rw [card_union_of_disjoint ‹_›] + convert lt_of_le_of_lt (add_le_add_left kk _) (add_lt_add_right size _) using 1 + convert Nat.choose_succ_succ _ _ using 3 + all_goals rwa [Nat.sub_one, Nat.succ_pred_eq_of_pos] + apply this.not_le + convert Set.Sized.card_le _ + · rw [Fintype.card_fin] + rw [coe_union, Set.sized_union] + refine ⟨‹_›, ?_⟩ + convert h𝒜bar.shadow_iterate + omega + end Finset From 4d43a68bfb071c919060f81c7032edaf16fe33e6 Mon Sep 17 00:00:00 2001 From: Jiang Jiedong Date: Mon, 28 Oct 2024 06:12:36 +0000 Subject: [PATCH 499/505] feat(FieldTheory): define induced action of subfields and intermediate fields (#18141) Add instances for induced actions of subfields and intermediate fields. Co-authored-by: Jiang Jiedong <107380768+jjdishere@users.noreply.github.com> --- Mathlib/Algebra/Field/Subfield.lean | 66 +++++++++++++++++++ .../FieldTheory/IntermediateField/Basic.lean | 58 ++++++++++++++++ 2 files changed, 124 insertions(+) diff --git a/Mathlib/Algebra/Field/Subfield.lean b/Mathlib/Algebra/Field/Subfield.lean index 694d6cf7e5504..c1d8534cddfbe 100644 --- a/Mathlib/Algebra/Field/Subfield.lean +++ b/Mathlib/Algebra/Field/Subfield.lean @@ -871,3 +871,69 @@ theorem comap_map (f : K →+* L) (s : Subfield K) : (s.map f).comap f = s := SetLike.coe_injective (Set.preimage_image_eq _ f.injective) end Subfield + +/-! ## Actions by `Subfield`s + +These are just copies of the definitions about `Subsemiring` starting from +`Subsemiring.MulAction`. +-/ +section Actions + +namespace Subfield + +variable {X Y} + +/-- The action by a subfield is the action by the underlying field. -/ +instance [SMul K X] (F : Subfield K) : SMul F X := + inferInstanceAs (SMul F.toSubsemiring X) + +theorem smul_def [SMul K X] {F : Subfield K} (g : F) (m : X) : g • m = (g : K) • m := + rfl + +instance smulCommClass_left [SMul K Y] [SMul X Y] [SMulCommClass K X Y] (F : Subfield K) : + SMulCommClass F X Y := + inferInstanceAs (SMulCommClass F.toSubsemiring X Y) + +instance smulCommClass_right [SMul X Y] [SMul K Y] [SMulCommClass X K Y] (F : Subfield K) : + SMulCommClass X F Y := + inferInstanceAs (SMulCommClass X F.toSubsemiring Y) + +/-- Note that this provides `IsScalarTower F K K` which is needed by `smul_mul_assoc`. -/ +instance [SMul X Y] [SMul K X] [SMul K Y] [IsScalarTower K X Y] (F : Subfield K) : + IsScalarTower F X Y := + inferInstanceAs (IsScalarTower F.toSubsemiring X Y) + +instance [SMul K X] [FaithfulSMul K X] (F : Subfield K) : FaithfulSMul F X := + inferInstanceAs (FaithfulSMul F.toSubsemiring X) + +/-- The action by a subfield is the action by the underlying field. -/ +instance [MulAction K X] (F : Subfield K) : MulAction F X := + inferInstanceAs (MulAction F.toSubsemiring X) + +/-- The action by a subfield is the action by the underlying field. -/ +instance [AddMonoid X] [DistribMulAction K X] (F : Subfield K) : DistribMulAction F X := + inferInstanceAs (DistribMulAction F.toSubsemiring X) + +/-- The action by a subfield is the action by the underlying field. -/ +instance [Monoid X] [MulDistribMulAction K X] (F : Subfield K) : MulDistribMulAction F X := + inferInstanceAs (MulDistribMulAction F.toSubsemiring X) + +/-- The action by a subfield is the action by the underlying field. -/ +instance [Zero X] [SMulWithZero K X] (F : Subfield K) : SMulWithZero F X := + inferInstanceAs (SMulWithZero F.toSubsemiring X) + +/-- The action by a subfield is the action by the underlying field. -/ +instance [Zero X] [MulActionWithZero K X] (F : Subfield K) : MulActionWithZero F X := + inferInstanceAs (MulActionWithZero F.toSubsemiring X) + +/-- The action by a subfield is the action by the underlying field. -/ +instance [AddCommMonoid X] [Module K X] (F : Subfield K) : Module F X := + inferInstanceAs (Module F.toSubsemiring X) + +/-- The action by a subfield is the action by the underlying field. -/ +instance [Semiring X] [MulSemiringAction K X] (F : Subfield K) : MulSemiringAction F X := + inferInstanceAs (MulSemiringAction F.toSubsemiring X) + +end Subfield + +end Actions diff --git a/Mathlib/FieldTheory/IntermediateField/Basic.lean b/Mathlib/FieldTheory/IntermediateField/Basic.lean index 23cb5fa01ffb4..581e1e65e3e36 100644 --- a/Mathlib/FieldTheory/IntermediateField/Basic.lean +++ b/Mathlib/FieldTheory/IntermediateField/Basic.lean @@ -302,6 +302,64 @@ theorem coe_prod {ι : Type*} [Fintype ι] (f : ι → S) : (↑(∏ i, f i) : L · simp · rw [Finset.prod_insert hi, MulMemClass.coe_mul, H, Finset.prod_insert hi] +/-! +`IntermediateField`s inherit structure from their `Subfield` coercions. +-/ + +variable {X Y} + +/-- The action by an intermediate field is the action by the underlying field. -/ +instance [SMul L X] (F : IntermediateField K L) : SMul F X := + inferInstanceAs (SMul F.toSubfield X) + +theorem smul_def [SMul L X] {F : IntermediateField K L} (g : F) (m : X) : g • m = (g : L) • m := + rfl + +instance smulCommClass_left [SMul L Y] [SMul X Y] [SMulCommClass L X Y] + (F : IntermediateField K L) : SMulCommClass F X Y := + inferInstanceAs (SMulCommClass F.toSubfield X Y) + +instance smulCommClass_right [SMul X Y] [SMul L Y] [SMulCommClass X L Y] + (F : IntermediateField K L) : SMulCommClass X F Y := + inferInstanceAs (SMulCommClass X F.toSubfield Y) + +/-- Note that this provides `IsScalarTower F K K` which is needed by `smul_mul_assoc`. -/ +instance [SMul X Y] [SMul L X] [SMul L Y] [IsScalarTower L X Y] (F : IntermediateField K L) : + IsScalarTower F X Y := + inferInstanceAs (IsScalarTower F.toSubfield X Y) + +instance [SMul L X] [FaithfulSMul L X] (F : IntermediateField K L) : FaithfulSMul F X := + inferInstanceAs (FaithfulSMul F.toSubfield X) + +/-- The action by an intermediate field is the action by the underlying field. -/ +instance [MulAction L X] (F : IntermediateField K L) : MulAction F X := + inferInstanceAs (MulAction F.toSubfield X) + +/-- The action by an intermediate field is the action by the underlying field. -/ +instance [AddMonoid X] [DistribMulAction L X] (F : IntermediateField K L) : DistribMulAction F X := + inferInstanceAs (DistribMulAction F.toSubfield X) + +/-- The action by an intermediate field is the action by the underlying field. -/ +instance [Monoid X] [MulDistribMulAction L X] (F : IntermediateField K L) : + MulDistribMulAction F X := + inferInstanceAs (MulDistribMulAction F.toSubfield X) + +/-- The action by an intermediate field is the action by the underlying field. -/ +instance [Zero X] [SMulWithZero L X] (F : IntermediateField K L) : SMulWithZero F X := + inferInstanceAs (SMulWithZero F.toSubfield X) + +/-- The action by an intermediate field is the action by the underlying field. -/ +instance [Zero X] [MulActionWithZero L X] (F : IntermediateField K L) : MulActionWithZero F X := + inferInstanceAs (MulActionWithZero F.toSubfield X) + +/-- The action by an intermediate field is the action by the underlying field. -/ +instance [AddCommMonoid X] [Module L X] (F : IntermediateField K L) : Module F X := + inferInstanceAs (Module F.toSubfield X) + +/-- The action by an intermediate field is the action by the underlying field. -/ +instance [Semiring X] [MulSemiringAction L X] (F : IntermediateField K L) : MulSemiringAction F X := + inferInstanceAs (MulSemiringAction F.toSubfield X) + /-! `IntermediateField`s inherit structure from their `Subalgebra` coercions. -/ instance toAlgebra : Algebra S L := From b8b4846e869616dcf97d496ab01cfef8acb3f3fe Mon Sep 17 00:00:00 2001 From: Jz Pan Date: Mon, 28 Oct 2024 06:12:37 +0000 Subject: [PATCH 500/505] feat(Algebra/Polynomial/AlgebraMap): add `algEquiv(OfCompEqX|AevalXAddC)_(eq_iff|symm)` (#18257) --- Mathlib/Algebra/Polynomial/AlgebraMap.lean | 20 ++++++++++++++++++++ 1 file changed, 20 insertions(+) diff --git a/Mathlib/Algebra/Polynomial/AlgebraMap.lean b/Mathlib/Algebra/Polynomial/AlgebraMap.lean index 1f0fbd0dedb84..5d623dcc85d62 100644 --- a/Mathlib/Algebra/Polynomial/AlgebraMap.lean +++ b/Mathlib/Algebra/Polynomial/AlgebraMap.lean @@ -294,12 +294,32 @@ def algEquivOfCompEqX (p q : R[X]) (hpq : p.comp q = X) (hqp : q.comp p = X) : R refine AlgEquiv.ofAlgHom (aeval p) (aeval q) ?_ ?_ <;> exact AlgHom.ext fun _ ↦ by simp [← comp_eq_aeval, comp_assoc, hpq, hqp] +@[simp] +theorem algEquivOfCompEqX_eq_iff (p q p' q' : R[X]) + (hpq : p.comp q = X) (hqp : q.comp p = X) (hpq' : p'.comp q' = X) (hqp' : q'.comp p' = X) : + algEquivOfCompEqX p q hpq hqp = algEquivOfCompEqX p' q' hpq' hqp' ↔ p = p' := + ⟨fun h ↦ by simpa using congr($h X), fun h ↦ by ext1; simp [h]⟩ + +@[simp] +theorem algEquivOfCompEqX_symm (p q : R[X]) (hpq : p.comp q = X) (hqp : q.comp p = X) : + (algEquivOfCompEqX p q hpq hqp).symm = algEquivOfCompEqX q p hqp hpq := rfl + /-- The automorphism of the polynomial algebra given by `p(X) ↦ p(X+t)`, with inverse `p(X) ↦ p(X-t)`. -/ @[simps!] def algEquivAevalXAddC {R} [CommRing R] (t : R) : R[X] ≃ₐ[R] R[X] := algEquivOfCompEqX (X + C t) (X - C t) (by simp) (by simp) +@[simp] +theorem algEquivAevalXAddC_eq_iff {R} [CommRing R] (t t' : R) : + algEquivAevalXAddC t = algEquivAevalXAddC t' ↔ t = t' := by + simp [algEquivAevalXAddC] + +@[simp] +theorem algEquivAevalXAddC_symm {R} [CommRing R] (t : R) : + (algEquivAevalXAddC t).symm = algEquivAevalXAddC (-t) := by + simp [algEquivAevalXAddC, sub_eq_add_neg] + theorem aeval_algHom (f : A →ₐ[R] B) (x : A) : aeval (f x) = f.comp (aeval x) := algHom_ext <| by simp only [aeval_X, AlgHom.comp_apply] From 95c5f3f542d052641f8129de7b204d115c687a76 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Mon, 28 Oct 2024 06:12:38 +0000 Subject: [PATCH 501/505] refactor: generalize AddSubmonoid Mul lemmas to SMul (#18278) --- .../Algebra/Group/Submonoid/Pointwise.lean | 146 ++++++++++-------- .../Algebra/GroupWithZero/Action/Defs.lean | 2 +- Mathlib/Algebra/Module/Defs.lean | 3 + 3 files changed, 88 insertions(+), 63 deletions(-) diff --git a/Mathlib/Algebra/Group/Submonoid/Pointwise.lean b/Mathlib/Algebra/Group/Submonoid/Pointwise.lean index 74f18f06313d8..fc97cb647e916 100644 --- a/Mathlib/Algebra/Group/Submonoid/Pointwise.lean +++ b/Mathlib/Algebra/Group/Submonoid/Pointwise.lean @@ -436,6 +436,69 @@ theorem one_eq_closure_one_set : (1 : AddSubmonoid R) = closure 1 := end AddMonoidWithOne +section SMul + +variable [AddMonoid R] [DistribSMul R A] + +/-- For `M : Submonoid R` and `N : AddSubmonoid A`, `M • N` is the additive submonoid +generated by all `m • n` where `m ∈ M` and `n ∈ N`. -/ +protected def smul : SMul (AddSubmonoid R) (AddSubmonoid A) where + smul M N := ⨆ s : M, N.map (DistribSMul.toAddMonoidHom A s.1) + +scoped[Pointwise] attribute [instance] AddSubmonoid.smul + +variable {M M' : AddSubmonoid R} {N P : AddSubmonoid A} {m : R} {n : A} + +theorem smul_mem_smul (hm : m ∈ M) (hn : n ∈ N) : m • n ∈ M • N := + (le_iSup _ ⟨m, hm⟩ : _ ≤ M • N) ⟨n, hn, by rfl⟩ + +theorem smul_le : M • N ≤ P ↔ ∀ m ∈ M, ∀ n ∈ N, m • n ∈ P := + ⟨fun H _m hm _n hn => H <| smul_mem_smul hm hn, fun H => + iSup_le fun ⟨m, hm⟩ => map_le_iff_le_comap.2 fun n hn => H m hm n hn⟩ + +@[elab_as_elim] +protected theorem smul_induction_on {C : A → Prop} {a : A} (ha : a ∈ M • N) + (hm : ∀ m ∈ M, ∀ n ∈ N, C (m • n)) (hadd : ∀ x y, C x → C y → C (x + y)) : C a := + (@smul_le _ _ _ _ _ _ _ ⟨⟨setOf C, hadd _ _⟩, by + simpa only [smul_zero] using hm _ (zero_mem _) _ (zero_mem _)⟩).2 hm ha + +@[simp] +theorem addSubmonoid_smul_bot (S : AddSubmonoid R) : S • (⊥ : AddSubmonoid A) = ⊥ := + eq_bot_iff.2 <| smul_le.2 fun m _ n hn => by + rw [AddSubmonoid.mem_bot] at hn ⊢; rw [hn, smul_zero] + +theorem smul_le_smul (h : M ≤ M') (hnp : N ≤ P) : M • N ≤ M' • P := + smul_le.2 fun _m hm _n hn => smul_mem_smul (h hm) (hnp hn) + +theorem smul_le_smul_left (h : M ≤ M') : M • P ≤ M' • P := + smul_le_smul h le_rfl + +theorem smul_le_smul_right (h : N ≤ P) : M • N ≤ M • P := + smul_le_smul le_rfl h + +theorem smul_subset_smul : (↑M : Set R) • (↑N : Set A) ⊆ (↑(M • N) : Set A) := + smul_subset_iff.2 fun _i hi _j hj ↦ smul_mem_smul hi hj + +theorem addSubmonoid_smul_sup : M • (N ⊔ P) = M • N ⊔ M • P := + le_antisymm (smul_le.mpr fun m hm np hnp ↦ by + refine closure_induction (p := (fun _ ↦ _ • · ∈ _)) ?_ ?_ ?_ (sup_eq_closure N P ▸ hnp) + · rintro x (hx | hx) + exacts [le_sup_left (a := M • N) (smul_mem_smul hm hx), + le_sup_right (a := M • N) (smul_mem_smul hm hx)] + · apply (smul_zero (A := A) m).symm ▸ (M • N ⊔ M • P).zero_mem + · intros _ _ _ _ h1 h2; rw [smul_add]; exact add_mem h1 h2) + (sup_le (smul_le_smul_right le_sup_left) <| smul_le_smul_right le_sup_right) + +variable {ι : Sort*} + +theorem smul_iSup (T : AddSubmonoid R) (S : ι → AddSubmonoid A) : (T • ⨆ i, S i) = ⨆ i, T • S i := + le_antisymm (smul_le.mpr fun t ht s hs ↦ iSup_induction _ (C := (t • · ∈ _)) hs + (fun i s hs ↦ mem_iSup_of_mem i <| smul_mem_smul ht hs) + (by simp_rw [smul_zero]; apply zero_mem) fun x y ↦ by simp_rw [smul_add]; apply add_mem) + (iSup_le fun i ↦ smul_le_smul_right <| le_iSup _ i) + +end SMul + section NonUnitalNonAssocSemiring variable [NonUnitalNonAssocSemiring R] @@ -447,20 +510,19 @@ protected def mul : Mul (AddSubmonoid R) := scoped[Pointwise] attribute [instance] AddSubmonoid.mul theorem mul_mem_mul {M N : AddSubmonoid R} {m n : R} (hm : m ∈ M) (hn : n ∈ N) : m * n ∈ M * N := - (le_iSup _ ⟨m, hm⟩ : _ ≤ M * N) ⟨n, hn, by rfl⟩ + smul_mem_smul hm hn theorem mul_le {M N P : AddSubmonoid R} : M * N ≤ P ↔ ∀ m ∈ M, ∀ n ∈ N, m * n ∈ P := - ⟨fun H _m hm _n hn => H <| mul_mem_mul hm hn, fun H => - iSup_le fun ⟨m, hm⟩ => map_le_iff_le_comap.2 fun n hn => H m hm n hn⟩ + smul_le @[elab_as_elim] protected theorem mul_induction_on {M N : AddSubmonoid R} {C : R → Prop} {r : R} (hr : r ∈ M * N) (hm : ∀ m ∈ M, ∀ n ∈ N, C (m * n)) (ha : ∀ x y, C x → C y → C (x + y)) : C r := - (@mul_le _ _ _ _ ⟨⟨setOf C, ha _ _⟩, by - simpa only [zero_mul] using hm _ (zero_mem _) _ (zero_mem _)⟩).2 hm hr + AddSubmonoid.smul_induction_on hr hm ha -- this proof is copied directly from `Submodule.span_mul_span` -- Porting note: proof rewritten +-- need `add_smul` to generalize to `SMul` theorem closure_mul_closure (S T : Set R) : closure S * closure T = closure (S * T) := by apply le_antisymm · refine mul_le.2 fun a ha b hb => ?_ @@ -480,9 +542,9 @@ theorem mul_eq_closure_mul_set (M N : AddSubmonoid R) : @[simp] theorem mul_bot (S : AddSubmonoid R) : S * ⊥ = ⊥ := - eq_bot_iff.2 <| mul_le.2 fun m _ n hn => by - rw [AddSubmonoid.mem_bot] at hn ⊢; rw [hn, mul_zero] + addSubmonoid_smul_bot S +-- need `zero_smul` to generalize to `SMul` @[simp] theorem bot_mul (S : AddSubmonoid R) : ⊥ * S = ⊥ := eq_bot_iff.2 <| mul_le.2 fun m hm n _ => by @@ -492,23 +554,21 @@ variable {M N P Q : AddSubmonoid R} @[mono] theorem mul_le_mul (hmp : M ≤ P) (hnq : N ≤ Q) : M * N ≤ P * Q := - mul_le.2 fun _m hm _n hn => mul_mem_mul (hmp hm) (hnq hn) + smul_le_smul hmp hnq theorem mul_le_mul_left (h : M ≤ N) : M * P ≤ N * P := - mul_le_mul h (le_refl P) + smul_le_smul_left h theorem mul_le_mul_right (h : N ≤ P) : M * N ≤ M * P := - mul_le_mul (le_refl M) h + smul_le_smul_right h theorem mul_subset_mul : (↑M : Set R) * (↑N : Set R) ⊆ (↑(M * N) : Set R) := - mul_subset_iff.2 fun _i hi _j hj ↦ mul_mem_mul hi hj + smul_subset_smul theorem mul_sup : M * (N ⊔ P) = M * N ⊔ M * P := - le_antisymm (mul_le.mpr fun m hm np hnp ↦ by - obtain ⟨n, hn, p, hp, rfl⟩ := mem_sup.mp hnp - rw [left_distrib]; exact add_mem_sup (mul_mem_mul hm hn) <| mul_mem_mul hm hp) - (sup_le (mul_le_mul_right le_sup_left) <| mul_le_mul_right le_sup_right) + addSubmonoid_smul_sup +-- need `zero_smul` and `add_smul` to generalize to `SMul` theorem sup_mul : (M ⊔ N) * P = M * P ⊔ N * P := le_antisymm (mul_le.mpr fun mn hmn p hp ↦ by obtain ⟨m, hm, n, hn, rfl⟩ := mem_sup.mp hmn @@ -517,6 +577,7 @@ theorem sup_mul : (M ⊔ N) * P = M * P ⊔ N * P := variable {ι : Sort*} +-- need `zero_smul` and `add_smul` to generalize to `SMul` theorem iSup_mul (S : ι → AddSubmonoid R) (T : AddSubmonoid R) : (⨆ i, S i) * T = ⨆ i, S i * T := le_antisymm (mul_le.mpr fun s hs t ht ↦ iSup_induction _ (C := (· * t ∈ _)) hs (fun i s hs ↦ mem_iSup_of_mem i <| mul_mem_mul hs ht) (by simp_rw [zero_mul]; apply zero_mem) @@ -524,10 +585,7 @@ theorem iSup_mul (S : ι → AddSubmonoid R) (T : AddSubmonoid R) : (⨆ i, S i) iSup_le fun i ↦ mul_le_mul_left (le_iSup _ i) theorem mul_iSup (T : AddSubmonoid R) (S : ι → AddSubmonoid R) : (T * ⨆ i, S i) = ⨆ i, T * S i := - le_antisymm (mul_le.mpr fun t ht s hs ↦ iSup_induction _ (C := (t * · ∈ _)) hs - (fun i s hs ↦ mem_iSup_of_mem i <| mul_mem_mul ht hs) (by simp_rw [mul_zero]; apply zero_mem) - fun _ _ ↦ by simp_rw [left_distrib]; apply add_mem) <| - iSup_le fun i ↦ mul_le_mul_right (le_iSup _ i) + smul_iSup T S end NonUnitalNonAssocSemiring @@ -580,17 +638,14 @@ variable [NonUnitalSemiring R] /-- Semigroup structure on additive submonoids of a (possibly, non-unital) semiring. -/ protected def semigroup : Semigroup (AddSubmonoid R) where mul := (· * ·) - mul_assoc M N P := + mul_assoc _M _N _P := le_antisymm - (mul_le.2 fun _mn hmn p hp => - suffices M * N ≤ (M * (N * P)).comap (AddMonoidHom.mulRight p) from this hmn - mul_le.2 fun m hm n hn => - show m * n * p ∈ M * (N * P) from - (mul_assoc m n p).symm ▸ mul_mem_mul hm (mul_mem_mul hn hp)) - (mul_le.2 fun m hm _np hnp => - suffices N * P ≤ (M * N * P).comap (AddMonoidHom.mulLeft m) from this hnp - mul_le.2 fun n hn p hp => - show m * (n * p) ∈ M * N * P from mul_assoc m n p ▸ mul_mem_mul (mul_mem_mul hm hn) hp) + (mul_le.2 fun _mn hmn p hp => AddSubmonoid.mul_induction_on hmn + (fun m hm n hn ↦ mul_assoc m n p ▸ mul_mem_mul hm <| mul_mem_mul hn hp) + fun x y ↦ (add_mul x y p).symm ▸ add_mem) + (mul_le.2 fun m hm _np hnp => AddSubmonoid.mul_induction_on hnp + (fun n hn p hp ↦ mul_assoc m n p ▸ mul_mem_mul (mul_mem_mul hm hn) hp) + fun x y ↦ (mul_add m x y) ▸ add_mem) scoped[Pointwise] attribute [instance] AddSubmonoid.semigroup end NonUnitalSemiring @@ -616,39 +671,6 @@ theorem pow_subset_pow {s : AddSubmonoid R} {n : ℕ} : (↑s : Set R) ^ n ⊆ end Semiring -section SMul - -variable [AddMonoid R] [DistribSMul R A] - -/-- For `M : Submonoid R` and `N : AddSubmonoid A`, `M • N` is the additive submonoid -generated by all `m • n` where `m ∈ M` and `n ∈ N`. -/ -protected def smul : SMul (AddSubmonoid R) (AddSubmonoid A) where - smul M N := ⨆ s : M, N.map (DistribSMul.toAddMonoidHom A s.1) - -scoped[Pointwise] attribute [instance] AddSubmonoid.smul - -example {R} [Semiring R] : Mul.toSMul (AddSubmonoid R) = AddSubmonoid.smul := rfl - -variable {M M' : AddSubmonoid R} {N P : AddSubmonoid A} {m : R} {n : A} - -theorem smul_mem_smul (hm : m ∈ M) (hn : n ∈ N) : m • n ∈ M • N := - (le_iSup _ ⟨m, hm⟩ : _ ≤ M • N) ⟨n, hn, by rfl⟩ - -theorem smul_le : M • N ≤ P ↔ ∀ m ∈ M, ∀ n ∈ N, m • n ∈ P := - ⟨fun H _m hm _n hn => H <| smul_mem_smul hm hn, fun H => - iSup_le fun ⟨m, hm⟩ => map_le_iff_le_comap.2 fun n hn => H m hm n hn⟩ - -@[elab_as_elim] -protected theorem smul_induction_on {C : A → Prop} {a : A} (ha : a ∈ M • N) - (hm : ∀ m ∈ M, ∀ n ∈ N, C (m • n)) (hadd : ∀ x y, C x → C y → C (x + y)) : C a := - (@smul_le _ _ _ _ _ _ _ ⟨⟨setOf C, hadd _ _⟩, by - simpa only [smul_zero] using hm _ (zero_mem _) _ (zero_mem _)⟩).2 hm ha - -theorem smul_le_smul (h : M ≤ M') (hnp : N ≤ P) : M • N ≤ M' • P := - smul_le.2 fun _m hm _n hn => smul_mem_smul (h hm) (hnp hn) - -end SMul - end AddSubmonoid namespace Set.IsPWO diff --git a/Mathlib/Algebra/GroupWithZero/Action/Defs.lean b/Mathlib/Algebra/GroupWithZero/Action/Defs.lean index 77eac698ee0d5..f4420acfd2b38 100644 --- a/Mathlib/Algebra/GroupWithZero/Action/Defs.lean +++ b/Mathlib/Algebra/GroupWithZero/Action/Defs.lean @@ -169,7 +169,7 @@ abbrev DistribSMul.compFun (f : N → M) : DistribSMul N A := /-- Each element of the scalars defines an additive monoid homomorphism. -/ @[simps] def DistribSMul.toAddMonoidHom (x : M) : A →+ A := - { SMulZeroClass.toZeroHom A x with toFun := (· • ·) x, map_add' := smul_add x } + { SMulZeroClass.toZeroHom A x with toFun := (x • ·), map_add' := smul_add x } end DistribSMul diff --git a/Mathlib/Algebra/Module/Defs.lean b/Mathlib/Algebra/Module/Defs.lean index 059efa2c28483..394dd7a778e54 100644 --- a/Mathlib/Algebra/Module/Defs.lean +++ b/Mathlib/Algebra/Module/Defs.lean @@ -240,6 +240,9 @@ instance (priority := 910) Semiring.toModule [Semiring R] : Module R R where zero_smul := zero_mul smul_zero := mul_zero +instance [NonUnitalNonAssocSemiring R] : DistribSMul R R where + smul_add := left_distrib + /-- A ring homomorphism `f : R →+* M` defines a module structure by `r • x = f r * x`. -/ def RingHom.toModule [Semiring R] [Semiring S] (f : R →+* S) : Module R S := Module.compHom S f From f4eb221ab8e0c96616b9e9fb490f864d5a2a8ab4 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 28 Oct 2024 06:12:39 +0000 Subject: [PATCH 502/505] feat: gluing measurable functions along measurable embeddings (#18287) --- .../MeasurableSpace/Embedding.lean | 32 +++++++++++++++++++ 1 file changed, 32 insertions(+) diff --git a/Mathlib/MeasureTheory/MeasurableSpace/Embedding.lean b/Mathlib/MeasureTheory/MeasurableSpace/Embedding.lean index 67ab6dc3c5625..6cd197d8d3f76 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/Embedding.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/Embedding.lean @@ -120,6 +120,38 @@ theorem measurable_comp_iff (hg : MeasurableEmbedding g) : Measurable (g ∘ f) end MeasurableEmbedding +section gluing +variable {α₁ α₂ α₃ : Type*} {mα : MeasurableSpace α} {mβ : MeasurableSpace β} + {mα₁ : MeasurableSpace α₁} {mα₂ : MeasurableSpace α₂} {mα₃ : MeasurableSpace α₃} + {i₁ : α₁ → α} {i₂ : α₂ → α} {i₃ : α₃ → α} {s : Set α} {f : α → β} + +lemma MeasurableSet.of_union_range_cover (hi₁ : MeasurableEmbedding i₁) + (hi₂ : MeasurableEmbedding i₂) (h : univ ⊆ range i₁ ∪ range i₂) + (hs₁ : MeasurableSet (i₁ ⁻¹' s)) (hs₂ : MeasurableSet (i₂ ⁻¹' s)) : MeasurableSet s := by + convert (hi₁.measurableSet_image' hs₁).union (hi₂.measurableSet_image' hs₂) + simp [image_preimage_eq_range_inter, ← union_inter_distrib_right,univ_subset_iff.1 h] + +lemma MeasurableSet.of_union₃_range_cover (hi₁ : MeasurableEmbedding i₁) + (hi₂ : MeasurableEmbedding i₂) (hi₃ : MeasurableEmbedding i₃) + (h : univ ⊆ range i₁ ∪ range i₂ ∪ range i₃) (hs₁ : MeasurableSet (i₁ ⁻¹' s)) + (hs₂ : MeasurableSet (i₂ ⁻¹' s)) (hs₃ : MeasurableSet (i₃ ⁻¹' s)) : MeasurableSet s := by + convert (hi₁.measurableSet_image' hs₁).union (hi₂.measurableSet_image' hs₂) |>.union + (hi₃.measurableSet_image' hs₃) + simp [image_preimage_eq_range_inter, ← union_inter_distrib_right, univ_subset_iff.1 h] + +lemma Measurable.of_union_range_cover (hi₁ : MeasurableEmbedding i₁) + (hi₂ : MeasurableEmbedding i₂) (h : univ ⊆ range i₁ ∪ range i₂) + (hf₁ : Measurable (f ∘ i₁)) (hf₂ : Measurable (f ∘ i₂)) : Measurable f := + fun _s hs ↦ .of_union_range_cover hi₁ hi₂ h (hf₁ hs) (hf₂ hs) + +lemma Measurable.of_union₃_range_cover (hi₁ : MeasurableEmbedding i₁) + (hi₂ : MeasurableEmbedding i₂) (hi₃ : MeasurableEmbedding i₃) + (h : univ ⊆ range i₁ ∪ range i₂ ∪ range i₃) (hf₁ : Measurable (f ∘ i₁)) + (hf₂ : Measurable (f ∘ i₂)) (hf₃ : Measurable (f ∘ i₃)) : Measurable f := + fun _s hs ↦ .of_union₃_range_cover hi₁ hi₂ hi₃ h (hf₁ hs) (hf₂ hs) (hf₃ hs) + +end gluing + theorem MeasurableSet.exists_measurable_proj {_ : MeasurableSpace α} (hs : MeasurableSet s) (hne : s.Nonempty) : ∃ f : α → s, Measurable f ∧ ∀ x : s, f x = x := let ⟨f, hfm, hf⟩ := From 091df8ec952a1cfe53fee8494ff1c917aecab921 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 28 Oct 2024 06:12:40 +0000 Subject: [PATCH 503/505] chore: split Mathlib.Topology.Algebra.UniformGroup (#18311) Another identified by #18156 --- Mathlib.lean | 3 +- Mathlib/Analysis/Convex/TotallyBounded.lean | 2 +- Mathlib/Analysis/LocallyConvex/Bounded.lean | 2 +- Mathlib/Analysis/Normed/Group/Uniform.lean | 2 +- Mathlib/Topology/Algebra/Equicontinuity.lean | 1 + Mathlib/Topology/Algebra/GroupCompletion.lean | 1 - .../Topology/Algebra/InfiniteSum/Group.lean | 3 +- Mathlib/Topology/Algebra/Module/Basic.lean | 3 +- .../Topology/Algebra/UniformConvergence.lean | 1 + .../Topology/Algebra/UniformFilterBasis.lean | 1 - .../Basic.lean} | 406 +------------ .../Topology/Algebra/UniformGroup/Defs.lean | 543 ++++++++++++++++++ .../Topology/Algebra/UniformMulAction.lean | 2 +- Mathlib/Topology/Algebra/UniformRing.lean | 1 + Mathlib/Topology/ContinuousMap/Algebra.lean | 1 - Mathlib/Topology/Instances/Real.lean | 1 + Mathlib/Topology/UniformSpace/Matrix.lean | 2 +- 17 files changed, 559 insertions(+), 416 deletions(-) rename Mathlib/Topology/Algebra/{UniformGroup.lean => UniformGroup/Basic.lean} (59%) create mode 100644 Mathlib/Topology/Algebra/UniformGroup/Defs.lean diff --git a/Mathlib.lean b/Mathlib.lean index e90ee6329b3f6..444f67315b74f 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4670,7 +4670,8 @@ import Mathlib.Topology.Algebra.StarSubalgebra import Mathlib.Topology.Algebra.UniformConvergence import Mathlib.Topology.Algebra.UniformField import Mathlib.Topology.Algebra.UniformFilterBasis -import Mathlib.Topology.Algebra.UniformGroup +import Mathlib.Topology.Algebra.UniformGroup.Basic +import Mathlib.Topology.Algebra.UniformGroup.Defs import Mathlib.Topology.Algebra.UniformMulAction import Mathlib.Topology.Algebra.UniformRing import Mathlib.Topology.Algebra.Valued.NormedValued diff --git a/Mathlib/Analysis/Convex/TotallyBounded.lean b/Mathlib/Analysis/Convex/TotallyBounded.lean index 5506f8e7dced5..f4215b59bc3cf 100644 --- a/Mathlib/Analysis/Convex/TotallyBounded.lean +++ b/Mathlib/Analysis/Convex/TotallyBounded.lean @@ -6,7 +6,7 @@ Authors: Christopher Hoskin import Mathlib.Topology.UniformSpace.Cauchy import Mathlib.Analysis.Convex.Hull -import Mathlib.Topology.Algebra.UniformGroup +import Mathlib.Topology.Algebra.UniformGroup.Basic import Mathlib.Topology.Algebra.Module.LocallyConvex /-! diff --git a/Mathlib/Analysis/LocallyConvex/Bounded.lean b/Mathlib/Analysis/LocallyConvex/Bounded.lean index e345035096ce2..1fbe648d3d553 100644 --- a/Mathlib/Analysis/LocallyConvex/Bounded.lean +++ b/Mathlib/Analysis/LocallyConvex/Bounded.lean @@ -9,7 +9,7 @@ import Mathlib.Analysis.LocallyConvex.BalancedCoreHull import Mathlib.Analysis.Seminorm import Mathlib.LinearAlgebra.Basis.VectorSpace import Mathlib.Topology.Bornology.Basic -import Mathlib.Topology.Algebra.UniformGroup +import Mathlib.Topology.Algebra.UniformGroup.Basic import Mathlib.Topology.UniformSpace.Cauchy import Mathlib.Topology.Algebra.Module.Basic diff --git a/Mathlib/Analysis/Normed/Group/Uniform.lean b/Mathlib/Analysis/Normed/Group/Uniform.lean index 3f63d8ab90c45..4cfd5c11b91c5 100644 --- a/Mathlib/Analysis/Normed/Group/Uniform.lean +++ b/Mathlib/Analysis/Normed/Group/Uniform.lean @@ -3,7 +3,7 @@ Copyright (c) 2018 Patrick Massot. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Patrick Massot, Johannes Hölzl, Yaël Dillies -/ -import Mathlib.Topology.Algebra.UniformGroup +import Mathlib.Topology.Algebra.UniformGroup.Basic import Mathlib.Topology.MetricSpace.Algebra import Mathlib.Topology.MetricSpace.IsometricSMul import Mathlib.Analysis.Normed.Group.Basic diff --git a/Mathlib/Topology/Algebra/Equicontinuity.lean b/Mathlib/Topology/Algebra/Equicontinuity.lean index d81fb26a3877a..b2eff285e0920 100644 --- a/Mathlib/Topology/Algebra/Equicontinuity.lean +++ b/Mathlib/Topology/Algebra/Equicontinuity.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Anatole Dedecker -/ import Mathlib.Topology.Algebra.UniformConvergence +import Mathlib.Topology.UniformSpace.Equicontinuity /-! # Algebra-related equicontinuity criterions diff --git a/Mathlib/Topology/Algebra/GroupCompletion.lean b/Mathlib/Topology/Algebra/GroupCompletion.lean index 5935b46b343e8..6cf8c9761c59d 100644 --- a/Mathlib/Topology/Algebra/GroupCompletion.lean +++ b/Mathlib/Topology/Algebra/GroupCompletion.lean @@ -3,7 +3,6 @@ Copyright (c) 2018 Patrick Massot. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Patrick Massot, Johannes Hölzl -/ -import Mathlib.Topology.Algebra.UniformGroup import Mathlib.Topology.Algebra.UniformMulAction import Mathlib.Topology.UniformSpace.Completion diff --git a/Mathlib/Topology/Algebra/InfiniteSum/Group.lean b/Mathlib/Topology/Algebra/InfiniteSum/Group.lean index c4029fed54fca..057b3e6ac56a0 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/Group.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/Group.lean @@ -5,7 +5,8 @@ Authors: Johannes Hölzl -/ import Mathlib.SetTheory.Cardinal.Finite import Mathlib.Topology.Algebra.InfiniteSum.Basic -import Mathlib.Topology.Algebra.UniformGroup +import Mathlib.Topology.UniformSpace.Cauchy +import Mathlib.Topology.Algebra.UniformGroup.Defs /-! # Infinite sums and products in topological groups diff --git a/Mathlib/Topology/Algebra/Module/Basic.lean b/Mathlib/Topology/Algebra/Module/Basic.lean index 70d2d820610b9..38e5f2704feec 100644 --- a/Mathlib/Topology/Algebra/Module/Basic.lean +++ b/Mathlib/Topology/Algebra/Module/Basic.lean @@ -11,8 +11,9 @@ import Mathlib.LinearAlgebra.Pi import Mathlib.LinearAlgebra.Projection import Mathlib.Topology.Algebra.MulAction import Mathlib.Topology.Algebra.Ring.Basic -import Mathlib.Topology.Algebra.UniformGroup import Mathlib.Topology.UniformSpace.UniformEmbedding +import Mathlib.Topology.Algebra.Group.Quotient +import Mathlib.Topology.Algebra.UniformGroup.Defs /-! # Theory of topological modules and continuous linear maps. diff --git a/Mathlib/Topology/Algebra/UniformConvergence.lean b/Mathlib/Topology/Algebra/UniformConvergence.lean index beee906368d6e..d734abc5df1d3 100644 --- a/Mathlib/Topology/Algebra/UniformConvergence.lean +++ b/Mathlib/Topology/Algebra/UniformConvergence.lean @@ -5,6 +5,7 @@ Authors: Anatole Dedecker -/ import Mathlib.Topology.Algebra.UniformMulAction import Mathlib.Algebra.Module.Pi +import Mathlib.Topology.UniformSpace.UniformConvergenceTopology /-! # Algebraic facts about the topology of uniform convergence diff --git a/Mathlib/Topology/Algebra/UniformFilterBasis.lean b/Mathlib/Topology/Algebra/UniformFilterBasis.lean index 09c3ba45192d2..055353936b31b 100644 --- a/Mathlib/Topology/Algebra/UniformFilterBasis.lean +++ b/Mathlib/Topology/Algebra/UniformFilterBasis.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Patrick Massot -/ import Mathlib.Topology.Algebra.FilterBasis -import Mathlib.Topology.Algebra.UniformGroup /-! # Uniform properties of neighborhood bases in topological algebra diff --git a/Mathlib/Topology/Algebra/UniformGroup.lean b/Mathlib/Topology/Algebra/UniformGroup/Basic.lean similarity index 59% rename from Mathlib/Topology/Algebra/UniformGroup.lean rename to Mathlib/Topology/Algebra/UniformGroup/Basic.lean index 379ce246257ff..9f937bc52f9d4 100644 --- a/Mathlib/Topology/Algebra/UniformGroup.lean +++ b/Mathlib/Topology/Algebra/UniformGroup/Basic.lean @@ -7,6 +7,7 @@ import Mathlib.Topology.UniformSpace.UniformConvergence import Mathlib.Topology.UniformSpace.UniformEmbedding import Mathlib.Topology.UniformSpace.CompleteSeparated import Mathlib.Topology.UniformSpace.Compact +import Mathlib.Topology.Algebra.UniformGroup.Defs import Mathlib.Topology.Algebra.Group.Quotient import Mathlib.Topology.DiscreteSubset import Mathlib.Tactic.Abel @@ -14,19 +15,8 @@ import Mathlib.Tactic.Abel /-! # Uniform structure on topological groups -This file defines uniform groups and its additive counterpart. These typeclasses should be -preferred over using `[TopologicalSpace α] [TopologicalGroup α]` since every topological -group naturally induces a uniform structure. - -## Main declarations -* `UniformGroup` and `UniformAddGroup`: Multiplicative and additive uniform groups, that - i.e., groups with uniformly continuous `(*)` and `(⁻¹)` / `(+)` and `(-)`. - ## Main results -* `TopologicalAddGroup.toUniformSpace` and `comm_topologicalAddGroup_is_uniform` can be used - to construct a canonical uniformity for a topological add group. - * extension of ℤ-bilinear maps to complete groups (useful for ring completions) * `QuotientGroup.completeSpace` and `QuotientAddGroup.completeSpace` guarantee that quotients @@ -44,122 +34,8 @@ open Filter Set variable {α : Type*} {β : Type*} -/-- A uniform group is a group in which multiplication and inversion are uniformly continuous. -/ -class UniformGroup (α : Type*) [UniformSpace α] [Group α] : Prop where - uniformContinuous_div : UniformContinuous fun p : α × α => p.1 / p.2 - -/-- A uniform additive group is an additive group in which addition - and negation are uniformly continuous. -/ -class UniformAddGroup (α : Type*) [UniformSpace α] [AddGroup α] : Prop where - uniformContinuous_sub : UniformContinuous fun p : α × α => p.1 - p.2 - -attribute [to_additive] UniformGroup - -@[to_additive] -theorem UniformGroup.mk' {α} [UniformSpace α] [Group α] - (h₁ : UniformContinuous fun p : α × α => p.1 * p.2) (h₂ : UniformContinuous fun p : α => p⁻¹) : - UniformGroup α := - ⟨by simpa only [div_eq_mul_inv] using - h₁.comp (uniformContinuous_fst.prod_mk (h₂.comp uniformContinuous_snd))⟩ - variable [UniformSpace α] [Group α] [UniformGroup α] -@[to_additive] -theorem uniformContinuous_div : UniformContinuous fun p : α × α => p.1 / p.2 := - UniformGroup.uniformContinuous_div - -@[to_additive] -theorem UniformContinuous.div [UniformSpace β] {f : β → α} {g : β → α} (hf : UniformContinuous f) - (hg : UniformContinuous g) : UniformContinuous fun x => f x / g x := - uniformContinuous_div.comp (hf.prod_mk hg) - -@[to_additive] -theorem UniformContinuous.inv [UniformSpace β] {f : β → α} (hf : UniformContinuous f) : - UniformContinuous fun x => (f x)⁻¹ := by - have : UniformContinuous fun x => 1 / f x := uniformContinuous_const.div hf - simp_all - -@[to_additive] -theorem uniformContinuous_inv : UniformContinuous fun x : α => x⁻¹ := - uniformContinuous_id.inv - -@[to_additive] -theorem UniformContinuous.mul [UniformSpace β] {f : β → α} {g : β → α} (hf : UniformContinuous f) - (hg : UniformContinuous g) : UniformContinuous fun x => f x * g x := by - have : UniformContinuous fun x => f x / (g x)⁻¹ := hf.div hg.inv - simp_all - -@[to_additive] -theorem uniformContinuous_mul : UniformContinuous fun p : α × α => p.1 * p.2 := - uniformContinuous_fst.mul uniformContinuous_snd - -@[to_additive] -theorem UniformContinuous.mul_const [UniformSpace β] {f : β → α} (hf : UniformContinuous f) - (a : α) : UniformContinuous fun x ↦ f x * a := - hf.mul uniformContinuous_const - -@[to_additive] -theorem UniformContinuous.const_mul [UniformSpace β] {f : β → α} (hf : UniformContinuous f) - (a : α) : UniformContinuous fun x ↦ a * f x := - uniformContinuous_const.mul hf - -@[to_additive] -theorem uniformContinuous_mul_left (a : α) : UniformContinuous fun b : α => a * b := - uniformContinuous_id.const_mul _ - -@[to_additive] -theorem uniformContinuous_mul_right (a : α) : UniformContinuous fun b : α => b * a := - uniformContinuous_id.mul_const _ - -@[to_additive] -theorem UniformContinuous.div_const [UniformSpace β] {f : β → α} (hf : UniformContinuous f) - (a : α) : UniformContinuous fun x ↦ f x / a := - hf.div uniformContinuous_const - -@[to_additive] -theorem uniformContinuous_div_const (a : α) : UniformContinuous fun b : α => b / a := - uniformContinuous_id.div_const _ - -@[to_additive UniformContinuous.const_nsmul] -theorem UniformContinuous.pow_const [UniformSpace β] {f : β → α} (hf : UniformContinuous f) : - ∀ n : ℕ, UniformContinuous fun x => f x ^ n - | 0 => by - simp_rw [pow_zero] - exact uniformContinuous_const - | n + 1 => by - simp_rw [pow_succ'] - exact hf.mul (hf.pow_const n) - -@[to_additive uniformContinuous_const_nsmul] -theorem uniformContinuous_pow_const (n : ℕ) : UniformContinuous fun x : α => x ^ n := - uniformContinuous_id.pow_const n - -@[to_additive UniformContinuous.const_zsmul] -theorem UniformContinuous.zpow_const [UniformSpace β] {f : β → α} (hf : UniformContinuous f) : - ∀ n : ℤ, UniformContinuous fun x => f x ^ n - | (n : ℕ) => by - simp_rw [zpow_natCast] - exact hf.pow_const _ - | Int.negSucc n => by - simp_rw [zpow_negSucc] - exact (hf.pow_const _).inv - -@[to_additive uniformContinuous_const_zsmul] -theorem uniformContinuous_zpow_const (n : ℤ) : UniformContinuous fun x : α => x ^ n := - uniformContinuous_id.zpow_const n - -@[to_additive] -instance (priority := 10) UniformGroup.to_topologicalGroup : TopologicalGroup α where - continuous_mul := uniformContinuous_mul.continuous - continuous_inv := uniformContinuous_inv.continuous - -@[to_additive] -instance [UniformSpace β] [Group β] [UniformGroup β] : UniformGroup (α × β) := - ⟨((uniformContinuous_fst.comp uniformContinuous_fst).div - (uniformContinuous_fst.comp uniformContinuous_snd)).prod_mk - ((uniformContinuous_snd.comp uniformContinuous_fst).div - (uniformContinuous_snd.comp uniformContinuous_snd))⟩ - @[to_additive] instance Pi.instUniformGroup {ι : Type*} {G : ι → Type*} [∀ i, UniformSpace (G i)] [∀ i, Group (G i)] [∀ i, UniformGroup (G i)] : UniformGroup (∀ i, G i) where @@ -167,17 +43,6 @@ instance Pi.instUniformGroup {ι : Type*} {G : ι → Type*} [∀ i, UniformSpac (uniformContinuous_proj G i).comp uniformContinuous_fst |>.div <| (uniformContinuous_proj G i).comp uniformContinuous_snd -@[to_additive] -theorem uniformity_translate_mul (a : α) : ((𝓤 α).map fun x : α × α => (x.1 * a, x.2 * a)) = 𝓤 α := - le_antisymm (uniformContinuous_id.mul uniformContinuous_const) - (calc - 𝓤 α = - ((𝓤 α).map fun x : α × α => (x.1 * a⁻¹, x.2 * a⁻¹)).map fun x : α × α => - (x.1 * a, x.2 * a) := by simp [Filter.map_map, Function.comp_def] - _ ≤ (𝓤 α).map fun x : α × α => (x.1 * a, x.2 * a) := - Filter.map_mono (uniformContinuous_id.mul uniformContinuous_const) - ) - @[to_additive] theorem isUniformEmbedding_translate_mul (a : α) : IsUniformEmbedding fun x : α => x * a := { comap_uniformity := by @@ -189,41 +54,10 @@ theorem isUniformEmbedding_translate_mul (a : α) : IsUniformEmbedding fun x : @[deprecated (since := "2024-10-01")] alias uniformEmbedding_translate_mul := isUniformEmbedding_translate_mul -namespace MulOpposite - -@[to_additive] -instance : UniformGroup αᵐᵒᵖ := - ⟨uniformContinuous_op.comp - ((uniformContinuous_unop.comp uniformContinuous_snd).inv.mul <| - uniformContinuous_unop.comp uniformContinuous_fst)⟩ - -end MulOpposite - section LatticeOps variable [Group β] -@[to_additive] -theorem uniformGroup_sInf {us : Set (UniformSpace β)} (h : ∀ u ∈ us, @UniformGroup β u _) : - @UniformGroup β (sInf us) _ := - -- Porting note: {_} does not find `sInf us` instance, see `continuousSMul_sInf` - @UniformGroup.mk β (_) _ <| - uniformContinuous_sInf_rng.mpr fun u hu => - uniformContinuous_sInf_dom₂ hu hu (@UniformGroup.uniformContinuous_div β u _ (h u hu)) - -@[to_additive] -theorem uniformGroup_iInf {ι : Sort*} {us' : ι → UniformSpace β} - (h' : ∀ i, @UniformGroup β (us' i) _) : @UniformGroup β (⨅ i, us' i) _ := by - rw [← sInf_range] - exact uniformGroup_sInf (Set.forall_mem_range.mpr h') - -@[to_additive] -theorem uniformGroup_inf {u₁ u₂ : UniformSpace β} (h₁ : @UniformGroup β u₁ _) - (h₂ : @UniformGroup β u₂ _) : @UniformGroup β (u₁ ⊓ u₂) _ := by - rw [inf_eq_iInf] - refine uniformGroup_iInf fun b => ?_ - cases b <;> assumption - @[to_additive] lemma IsUniformInducing.uniformGroup {γ : Type*} [Group γ] [UniformSpace γ] [UniformGroup γ] [UniformSpace β] {F : Type*} [FunLike F β γ] [MonoidHomClass F β γ] @@ -250,147 +84,6 @@ instance uniformGroup (S : Subgroup α) : UniformGroup S := .comap S.subtype end Subgroup -section - -variable (α) - -@[to_additive] -theorem uniformity_eq_comap_nhds_one : 𝓤 α = comap (fun x : α × α => x.2 / x.1) (𝓝 (1 : α)) := by - rw [nhds_eq_comap_uniformity, Filter.comap_comap] - refine le_antisymm (Filter.map_le_iff_le_comap.1 ?_) ?_ - · intro s hs - rcases mem_uniformity_of_uniformContinuous_invariant uniformContinuous_div hs with ⟨t, ht, hts⟩ - refine mem_map.2 (mem_of_superset ht ?_) - rintro ⟨a, b⟩ - simpa [subset_def] using hts a b a - · intro s hs - rcases mem_uniformity_of_uniformContinuous_invariant uniformContinuous_mul hs with ⟨t, ht, hts⟩ - refine ⟨_, ht, ?_⟩ - rintro ⟨a, b⟩ - simpa [subset_def] using hts 1 (b / a) a - -@[to_additive] -theorem uniformity_eq_comap_nhds_one_swapped : - 𝓤 α = comap (fun x : α × α => x.1 / x.2) (𝓝 (1 : α)) := by - rw [← comap_swap_uniformity, uniformity_eq_comap_nhds_one, comap_comap] - rfl - -@[to_additive] -theorem UniformGroup.ext {G : Type*} [Group G] {u v : UniformSpace G} (hu : @UniformGroup G u _) - (hv : @UniformGroup G v _) - (h : @nhds _ u.toTopologicalSpace 1 = @nhds _ v.toTopologicalSpace 1) : u = v := - UniformSpace.ext <| by - rw [@uniformity_eq_comap_nhds_one _ u _ hu, @uniformity_eq_comap_nhds_one _ v _ hv, h] - -@[to_additive] -theorem UniformGroup.ext_iff {G : Type*} [Group G] {u v : UniformSpace G} - (hu : @UniformGroup G u _) (hv : @UniformGroup G v _) : - u = v ↔ @nhds _ u.toTopologicalSpace 1 = @nhds _ v.toTopologicalSpace 1 := - ⟨fun h => h ▸ rfl, hu.ext hv⟩ - -variable {α} - -@[to_additive] -theorem UniformGroup.uniformity_countably_generated [(𝓝 (1 : α)).IsCountablyGenerated] : - (𝓤 α).IsCountablyGenerated := by - rw [uniformity_eq_comap_nhds_one] - exact Filter.comap.isCountablyGenerated _ _ - -open MulOpposite - -@[to_additive] -theorem uniformity_eq_comap_inv_mul_nhds_one : - 𝓤 α = comap (fun x : α × α => x.1⁻¹ * x.2) (𝓝 (1 : α)) := by - rw [← comap_uniformity_mulOpposite, uniformity_eq_comap_nhds_one, ← op_one, ← comap_unop_nhds, - comap_comap, comap_comap] - simp [Function.comp_def] - -@[to_additive] -theorem uniformity_eq_comap_inv_mul_nhds_one_swapped : - 𝓤 α = comap (fun x : α × α => x.2⁻¹ * x.1) (𝓝 (1 : α)) := by - rw [← comap_swap_uniformity, uniformity_eq_comap_inv_mul_nhds_one, comap_comap] - rfl - -end - -@[to_additive] -theorem Filter.HasBasis.uniformity_of_nhds_one {ι} {p : ι → Prop} {U : ι → Set α} - (h : (𝓝 (1 : α)).HasBasis p U) : - (𝓤 α).HasBasis p fun i => { x : α × α | x.2 / x.1 ∈ U i } := by - rw [uniformity_eq_comap_nhds_one] - exact h.comap _ - -@[to_additive] -theorem Filter.HasBasis.uniformity_of_nhds_one_inv_mul {ι} {p : ι → Prop} {U : ι → Set α} - (h : (𝓝 (1 : α)).HasBasis p U) : - (𝓤 α).HasBasis p fun i => { x : α × α | x.1⁻¹ * x.2 ∈ U i } := by - rw [uniformity_eq_comap_inv_mul_nhds_one] - exact h.comap _ - -@[to_additive] -theorem Filter.HasBasis.uniformity_of_nhds_one_swapped {ι} {p : ι → Prop} {U : ι → Set α} - (h : (𝓝 (1 : α)).HasBasis p U) : - (𝓤 α).HasBasis p fun i => { x : α × α | x.1 / x.2 ∈ U i } := by - rw [uniformity_eq_comap_nhds_one_swapped] - exact h.comap _ - -@[to_additive] -theorem Filter.HasBasis.uniformity_of_nhds_one_inv_mul_swapped {ι} {p : ι → Prop} {U : ι → Set α} - (h : (𝓝 (1 : α)).HasBasis p U) : - (𝓤 α).HasBasis p fun i => { x : α × α | x.2⁻¹ * x.1 ∈ U i } := by - rw [uniformity_eq_comap_inv_mul_nhds_one_swapped] - exact h.comap _ - -@[to_additive] -theorem uniformContinuous_of_tendsto_one {hom : Type*} [UniformSpace β] [Group β] [UniformGroup β] - [FunLike hom α β] [MonoidHomClass hom α β] {f : hom} (h : Tendsto f (𝓝 1) (𝓝 1)) : - UniformContinuous f := by - have : - ((fun x : β × β => x.2 / x.1) ∘ fun x : α × α => (f x.1, f x.2)) = fun x : α × α => - f (x.2 / x.1) := by ext; simp only [Function.comp_apply, map_div] - rw [UniformContinuous, uniformity_eq_comap_nhds_one α, uniformity_eq_comap_nhds_one β, - tendsto_comap_iff, this] - exact Tendsto.comp h tendsto_comap - -/-- A group homomorphism (a bundled morphism of a type that implements `MonoidHomClass`) between -two uniform groups is uniformly continuous provided that it is continuous at one. See also -`continuous_of_continuousAt_one`. -/ -@[to_additive "An additive group homomorphism (a bundled morphism of a type that implements -`AddMonoidHomClass`) between two uniform additive groups is uniformly continuous provided that it -is continuous at zero. See also `continuous_of_continuousAt_zero`."] -theorem uniformContinuous_of_continuousAt_one {hom : Type*} [UniformSpace β] [Group β] - [UniformGroup β] [FunLike hom α β] [MonoidHomClass hom α β] - (f : hom) (hf : ContinuousAt f 1) : - UniformContinuous f := - uniformContinuous_of_tendsto_one (by simpa using hf.tendsto) - -@[to_additive] -theorem MonoidHom.uniformContinuous_of_continuousAt_one [UniformSpace β] [Group β] [UniformGroup β] - (f : α →* β) (hf : ContinuousAt f 1) : UniformContinuous f := - _root_.uniformContinuous_of_continuousAt_one f hf - -/-- A homomorphism from a uniform group to a discrete uniform group is continuous if and only if -its kernel is open. -/ -@[to_additive "A homomorphism from a uniform additive group to a discrete uniform additive group is -continuous if and only if its kernel is open."] -theorem UniformGroup.uniformContinuous_iff_open_ker {hom : Type*} [UniformSpace β] - [DiscreteTopology β] [Group β] [UniformGroup β] [FunLike hom α β] [MonoidHomClass hom α β] - {f : hom} : - UniformContinuous f ↔ IsOpen ((f : α →* β).ker : Set α) := by - refine ⟨fun hf => ?_, fun hf => ?_⟩ - · apply (isOpen_discrete ({1} : Set β)).preimage hf.continuous - · apply uniformContinuous_of_continuousAt_one - rw [ContinuousAt, nhds_discrete β, map_one, tendsto_pure] - exact hf.mem_nhds (map_one f) - -@[to_additive] -theorem uniformContinuous_monoidHom_of_continuous {hom : Type*} [UniformSpace β] [Group β] - [UniformGroup β] [FunLike hom α β] [MonoidHomClass hom α β] {f : hom} (h : Continuous f) : - UniformContinuous f := - uniformContinuous_of_tendsto_one <| - suffices Tendsto f (𝓝 1) (𝓝 (f 1)) by rwa [map_one] at this - h.tendsto 1 - @[to_additive] theorem CauchySeq.mul {ι : Type*} [Preorder ι] {u v : ι → α} (hu : CauchySeq u) (hv : CauchySeq v) : CauchySeq (u * v) := @@ -478,38 +171,8 @@ open Filter variable (G : Type*) [Group G] [TopologicalSpace G] [TopologicalGroup G] -/-- The right uniformity on a topological group (as opposed to the left uniformity). - -Warning: in general the right and left uniformities do not coincide and so one does not obtain a -`UniformGroup` structure. Two important special cases where they _do_ coincide are for -commutative groups (see `comm_topologicalGroup_is_uniform`) and for compact groups (see -`topologicalGroup_is_uniform_of_compactSpace`). -/ -@[to_additive "The right uniformity on a topological additive group (as opposed to the left -uniformity). - -Warning: in general the right and left uniformities do not coincide and so one does not obtain a -`UniformAddGroup` structure. Two important special cases where they _do_ coincide are for -commutative additive groups (see `comm_topologicalAddGroup_is_uniform`) and for compact -additive groups (see `topologicalAddGroup_is_uniform_of_compactSpace`)."] -def TopologicalGroup.toUniformSpace : UniformSpace G where - uniformity := comap (fun p : G × G => p.2 / p.1) (𝓝 1) - symm := - have : Tendsto (fun p : G × G ↦ (p.2 / p.1)⁻¹) (comap (fun p : G × G ↦ p.2 / p.1) (𝓝 1)) - (𝓝 1⁻¹) := tendsto_id.inv.comp tendsto_comap - by simpa [tendsto_comap_iff] - comp := Tendsto.le_comap fun U H ↦ by - rcases exists_nhds_one_split H with ⟨V, V_nhds, V_mul⟩ - refine mem_map.2 (mem_of_superset (mem_lift' <| preimage_mem_comap V_nhds) ?_) - rintro ⟨x, y⟩ ⟨z, hz₁, hz₂⟩ - simpa using V_mul _ hz₂ _ hz₁ - nhds_eq_comap_uniformity _ := by simp only [comap_comap, Function.comp_def, nhds_translation_div] - attribute [local instance] TopologicalGroup.toUniformSpace -@[to_additive] -theorem uniformity_eq_comap_nhds_one' : 𝓤 G = comap (fun p : G × G => p.2 / p.1) (𝓝 (1 : G)) := - rfl - @[to_additive] theorem topologicalGroup_is_uniform_of_compactSpace [CompactSpace G] : UniformGroup G := ⟨by @@ -580,75 +243,8 @@ theorem TopologicalGroup.tendstoLocallyUniformlyOn_iff {ι α : Type*} [Topologi end TopologicalGroup -section TopologicalCommGroup - -universe u v w x - -open Filter - -variable (G : Type*) [CommGroup G] [TopologicalSpace G] [TopologicalGroup G] - -section - -attribute [local instance] TopologicalGroup.toUniformSpace - -variable {G} - -@[to_additive] --- Porting note: renamed theorem to conform to naming convention -theorem comm_topologicalGroup_is_uniform : UniformGroup G := by - have : - Tendsto - ((fun p : G × G => p.1 / p.2) ∘ fun p : (G × G) × G × G => (p.1.2 / p.1.1, p.2.2 / p.2.1)) - (comap (fun p : (G × G) × G × G => (p.1.2 / p.1.1, p.2.2 / p.2.1)) ((𝓝 1).prod (𝓝 1))) - (𝓝 (1 / 1)) := - (tendsto_fst.div' tendsto_snd).comp tendsto_comap - constructor - rw [UniformContinuous, uniformity_prod_eq_prod, tendsto_map'_iff, uniformity_eq_comap_nhds_one' G, - tendsto_comap_iff, prod_comap_comap_eq] - simp only [Function.comp_def, div_eq_mul_inv, mul_inv_rev, inv_inv, mul_comm, mul_left_comm] at * - simp only [inv_one, mul_one, ← mul_assoc] at this - simp_rw [← mul_assoc, mul_comm] - assumption - -open Set - -end - -@[to_additive] -theorem UniformGroup.toUniformSpace_eq {G : Type*} [u : UniformSpace G] [Group G] - [UniformGroup G] : TopologicalGroup.toUniformSpace G = u := by - ext : 1 - rw [uniformity_eq_comap_nhds_one' G, uniformity_eq_comap_nhds_one G] - -end TopologicalCommGroup - open Filter Set Function -section - -variable {α : Type*} {β : Type*} {hom : Type*} -variable [TopologicalSpace α] [Group α] [TopologicalGroup α] - --- β is a dense subgroup of α, inclusion is denoted by e -variable [TopologicalSpace β] [Group β] -variable [FunLike hom β α] [MonoidHomClass hom β α] {e : hom} - -@[to_additive] -theorem tendsto_div_comap_self (de : IsDenseInducing e) (x₀ : α) : - Tendsto (fun t : β × β => t.2 / t.1) ((comap fun p : β × β => (e p.1, e p.2)) <| 𝓝 (x₀, x₀)) - (𝓝 1) := by - have comm : ((fun x : α × α => x.2 / x.1) ∘ fun t : β × β => (e t.1, e t.2)) = - e ∘ fun t : β × β => t.2 / t.1 := by - ext t - change e t.2 / e t.1 = e (t.2 / t.1) - rw [← map_div e t.2 t.1] - have lim : Tendsto (fun x : α × α => x.2 / x.1) (𝓝 (x₀, x₀)) (𝓝 (e 1)) := by - simpa using (continuous_div'.comp (@continuous_swap α α _ _)).tendsto (x₀, x₀) - simpa using de.tendsto_comap_nhds_nhds lim comm - -end - namespace IsDenseInducing variable {α : Type*} {β : Type*} {γ : Type*} {δ : Type*} diff --git a/Mathlib/Topology/Algebra/UniformGroup/Defs.lean b/Mathlib/Topology/Algebra/UniformGroup/Defs.lean new file mode 100644 index 0000000000000..1ebc4e5aa7b0d --- /dev/null +++ b/Mathlib/Topology/Algebra/UniformGroup/Defs.lean @@ -0,0 +1,543 @@ +/- +Copyright (c) 2018 Patrick Massot. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Patrick Massot, Johannes Hölzl +-/ +import Mathlib.Topology.UniformSpace.Basic +import Mathlib.Topology.Algebra.Group.Basic + +/-! +# Uniform structure on topological groups + +This file defines uniform groups and its additive counterpart. These typeclasses should be +preferred over using `[TopologicalSpace α] [TopologicalGroup α]` since every topological +group naturally induces a uniform structure. + +## Main declarations +* `UniformGroup` and `UniformAddGroup`: Multiplicative and additive uniform groups, + i.e., groups with uniformly continuous `(*)` and `(⁻¹)` / `(+)` and `(-)`. + +## Main results + +* `TopologicalAddGroup.toUniformSpace` and `comm_topologicalAddGroup_is_uniform` can be used + to construct a canonical uniformity for a topological add group. + +See `Mathlib.Topology.Algebra.UniformGroup.Basic` for further results. +-/ + +assert_not_exists Cauchy + +noncomputable section + +open Uniformity Topology Filter Pointwise + +section UniformGroup + +open Filter Set + +variable {α : Type*} {β : Type*} + +/-- A uniform group is a group in which multiplication and inversion are uniformly continuous. -/ +class UniformGroup (α : Type*) [UniformSpace α] [Group α] : Prop where + uniformContinuous_div : UniformContinuous fun p : α × α => p.1 / p.2 + +/-- A uniform additive group is an additive group in which addition + and negation are uniformly continuous. -/ +class UniformAddGroup (α : Type*) [UniformSpace α] [AddGroup α] : Prop where + uniformContinuous_sub : UniformContinuous fun p : α × α => p.1 - p.2 + +attribute [to_additive] UniformGroup + +@[to_additive] +theorem UniformGroup.mk' {α} [UniformSpace α] [Group α] + (h₁ : UniformContinuous fun p : α × α => p.1 * p.2) (h₂ : UniformContinuous fun p : α => p⁻¹) : + UniformGroup α := + ⟨by simpa only [div_eq_mul_inv] using + h₁.comp (uniformContinuous_fst.prod_mk (h₂.comp uniformContinuous_snd))⟩ + +variable [UniformSpace α] [Group α] [UniformGroup α] + +@[to_additive] +theorem uniformContinuous_div : UniformContinuous fun p : α × α => p.1 / p.2 := + UniformGroup.uniformContinuous_div + +@[to_additive] +theorem UniformContinuous.div [UniformSpace β] {f : β → α} {g : β → α} (hf : UniformContinuous f) + (hg : UniformContinuous g) : UniformContinuous fun x => f x / g x := + uniformContinuous_div.comp (hf.prod_mk hg) + +@[to_additive] +theorem UniformContinuous.inv [UniformSpace β] {f : β → α} (hf : UniformContinuous f) : + UniformContinuous fun x => (f x)⁻¹ := by + have : UniformContinuous fun x => 1 / f x := uniformContinuous_const.div hf + simp_all + +@[to_additive] +theorem uniformContinuous_inv : UniformContinuous fun x : α => x⁻¹ := + uniformContinuous_id.inv + +@[to_additive] +theorem UniformContinuous.mul [UniformSpace β] {f : β → α} {g : β → α} (hf : UniformContinuous f) + (hg : UniformContinuous g) : UniformContinuous fun x => f x * g x := by + have : UniformContinuous fun x => f x / (g x)⁻¹ := hf.div hg.inv + simp_all + +@[to_additive] +theorem uniformContinuous_mul : UniformContinuous fun p : α × α => p.1 * p.2 := + uniformContinuous_fst.mul uniformContinuous_snd + +@[to_additive] +theorem UniformContinuous.mul_const [UniformSpace β] {f : β → α} (hf : UniformContinuous f) + (a : α) : UniformContinuous fun x ↦ f x * a := + hf.mul uniformContinuous_const + +@[to_additive] +theorem UniformContinuous.const_mul [UniformSpace β] {f : β → α} (hf : UniformContinuous f) + (a : α) : UniformContinuous fun x ↦ a * f x := + uniformContinuous_const.mul hf + +@[to_additive] +theorem uniformContinuous_mul_left (a : α) : UniformContinuous fun b : α => a * b := + uniformContinuous_id.const_mul _ + +@[to_additive] +theorem uniformContinuous_mul_right (a : α) : UniformContinuous fun b : α => b * a := + uniformContinuous_id.mul_const _ + +@[to_additive] +theorem UniformContinuous.div_const [UniformSpace β] {f : β → α} (hf : UniformContinuous f) + (a : α) : UniformContinuous fun x ↦ f x / a := + hf.div uniformContinuous_const + +@[to_additive] +theorem uniformContinuous_div_const (a : α) : UniformContinuous fun b : α => b / a := + uniformContinuous_id.div_const _ + +@[to_additive UniformContinuous.const_nsmul] +theorem UniformContinuous.pow_const [UniformSpace β] {f : β → α} (hf : UniformContinuous f) : + ∀ n : ℕ, UniformContinuous fun x => f x ^ n + | 0 => by + simp_rw [pow_zero] + exact uniformContinuous_const + | n + 1 => by + simp_rw [pow_succ'] + exact hf.mul (hf.pow_const n) + +@[to_additive uniformContinuous_const_nsmul] +theorem uniformContinuous_pow_const (n : ℕ) : UniformContinuous fun x : α => x ^ n := + uniformContinuous_id.pow_const n + +@[to_additive UniformContinuous.const_zsmul] +theorem UniformContinuous.zpow_const [UniformSpace β] {f : β → α} (hf : UniformContinuous f) : + ∀ n : ℤ, UniformContinuous fun x => f x ^ n + | (n : ℕ) => by + simp_rw [zpow_natCast] + exact hf.pow_const _ + | Int.negSucc n => by + simp_rw [zpow_negSucc] + exact (hf.pow_const _).inv + +@[to_additive uniformContinuous_const_zsmul] +theorem uniformContinuous_zpow_const (n : ℤ) : UniformContinuous fun x : α => x ^ n := + uniformContinuous_id.zpow_const n + +@[to_additive] +instance (priority := 10) UniformGroup.to_topologicalGroup : TopologicalGroup α where + continuous_mul := uniformContinuous_mul.continuous + continuous_inv := uniformContinuous_inv.continuous + +@[to_additive] +instance [UniformSpace β] [Group β] [UniformGroup β] : UniformGroup (α × β) := + ⟨((uniformContinuous_fst.comp uniformContinuous_fst).div + (uniformContinuous_fst.comp uniformContinuous_snd)).prod_mk + ((uniformContinuous_snd.comp uniformContinuous_fst).div + (uniformContinuous_snd.comp uniformContinuous_snd))⟩ + +@[to_additive] +theorem uniformity_translate_mul (a : α) : ((𝓤 α).map fun x : α × α => (x.1 * a, x.2 * a)) = 𝓤 α := + le_antisymm (uniformContinuous_id.mul uniformContinuous_const) + (calc + 𝓤 α = + ((𝓤 α).map fun x : α × α => (x.1 * a⁻¹, x.2 * a⁻¹)).map fun x : α × α => + (x.1 * a, x.2 * a) := by simp [Filter.map_map, Function.comp_def] + _ ≤ (𝓤 α).map fun x : α × α => (x.1 * a, x.2 * a) := + Filter.map_mono (uniformContinuous_id.mul uniformContinuous_const) + ) + +namespace MulOpposite + +@[to_additive] +instance : UniformGroup αᵐᵒᵖ := + ⟨uniformContinuous_op.comp + ((uniformContinuous_unop.comp uniformContinuous_snd).inv.mul <| + uniformContinuous_unop.comp uniformContinuous_fst)⟩ + +end MulOpposite + +section LatticeOps + +variable [Group β] + +@[to_additive] +theorem uniformGroup_sInf {us : Set (UniformSpace β)} (h : ∀ u ∈ us, @UniformGroup β u _) : + @UniformGroup β (sInf us) _ := + -- Porting note: {_} does not find `sInf us` instance, see `continuousSMul_sInf` + @UniformGroup.mk β (_) _ <| + uniformContinuous_sInf_rng.mpr fun u hu => + uniformContinuous_sInf_dom₂ hu hu (@UniformGroup.uniformContinuous_div β u _ (h u hu)) + +@[to_additive] +theorem uniformGroup_iInf {ι : Sort*} {us' : ι → UniformSpace β} + (h' : ∀ i, @UniformGroup β (us' i) _) : @UniformGroup β (⨅ i, us' i) _ := by + rw [← sInf_range] + exact uniformGroup_sInf (Set.forall_mem_range.mpr h') + +@[to_additive] +theorem uniformGroup_inf {u₁ u₂ : UniformSpace β} (h₁ : @UniformGroup β u₁ _) + (h₂ : @UniformGroup β u₂ _) : @UniformGroup β (u₁ ⊓ u₂) _ := by + rw [inf_eq_iInf] + refine uniformGroup_iInf fun b => ?_ + cases b <;> assumption + +end LatticeOps + +section + +variable (α) + +@[to_additive] +theorem uniformity_eq_comap_nhds_one : 𝓤 α = comap (fun x : α × α => x.2 / x.1) (𝓝 (1 : α)) := by + rw [nhds_eq_comap_uniformity, Filter.comap_comap] + refine le_antisymm (Filter.map_le_iff_le_comap.1 ?_) ?_ + · intro s hs + rcases mem_uniformity_of_uniformContinuous_invariant uniformContinuous_div hs with ⟨t, ht, hts⟩ + refine mem_map.2 (mem_of_superset ht ?_) + rintro ⟨a, b⟩ + simpa [subset_def] using hts a b a + · intro s hs + rcases mem_uniformity_of_uniformContinuous_invariant uniformContinuous_mul hs with ⟨t, ht, hts⟩ + refine ⟨_, ht, ?_⟩ + rintro ⟨a, b⟩ + simpa [subset_def] using hts 1 (b / a) a + +@[to_additive] +theorem uniformity_eq_comap_nhds_one_swapped : + 𝓤 α = comap (fun x : α × α => x.1 / x.2) (𝓝 (1 : α)) := by + rw [← comap_swap_uniformity, uniformity_eq_comap_nhds_one, comap_comap] + rfl + +@[to_additive] +theorem UniformGroup.ext {G : Type*} [Group G] {u v : UniformSpace G} (hu : @UniformGroup G u _) + (hv : @UniformGroup G v _) + (h : @nhds _ u.toTopologicalSpace 1 = @nhds _ v.toTopologicalSpace 1) : u = v := + UniformSpace.ext <| by + rw [@uniformity_eq_comap_nhds_one _ u _ hu, @uniformity_eq_comap_nhds_one _ v _ hv, h] + +@[to_additive] +theorem UniformGroup.ext_iff {G : Type*} [Group G] {u v : UniformSpace G} + (hu : @UniformGroup G u _) (hv : @UniformGroup G v _) : + u = v ↔ @nhds _ u.toTopologicalSpace 1 = @nhds _ v.toTopologicalSpace 1 := + ⟨fun h => h ▸ rfl, hu.ext hv⟩ + +variable {α} + +@[to_additive] +theorem UniformGroup.uniformity_countably_generated [(𝓝 (1 : α)).IsCountablyGenerated] : + (𝓤 α).IsCountablyGenerated := by + rw [uniformity_eq_comap_nhds_one] + exact Filter.comap.isCountablyGenerated _ _ + +open MulOpposite + +@[to_additive] +theorem uniformity_eq_comap_inv_mul_nhds_one : + 𝓤 α = comap (fun x : α × α => x.1⁻¹ * x.2) (𝓝 (1 : α)) := by + rw [← comap_uniformity_mulOpposite, uniformity_eq_comap_nhds_one, ← op_one, ← comap_unop_nhds, + comap_comap, comap_comap] + simp [Function.comp_def] + +@[to_additive] +theorem uniformity_eq_comap_inv_mul_nhds_one_swapped : + 𝓤 α = comap (fun x : α × α => x.2⁻¹ * x.1) (𝓝 (1 : α)) := by + rw [← comap_swap_uniformity, uniformity_eq_comap_inv_mul_nhds_one, comap_comap] + rfl + +end + +@[to_additive] +theorem Filter.HasBasis.uniformity_of_nhds_one {ι} {p : ι → Prop} {U : ι → Set α} + (h : (𝓝 (1 : α)).HasBasis p U) : + (𝓤 α).HasBasis p fun i => { x : α × α | x.2 / x.1 ∈ U i } := by + rw [uniformity_eq_comap_nhds_one] + exact h.comap _ + +@[to_additive] +theorem Filter.HasBasis.uniformity_of_nhds_one_inv_mul {ι} {p : ι → Prop} {U : ι → Set α} + (h : (𝓝 (1 : α)).HasBasis p U) : + (𝓤 α).HasBasis p fun i => { x : α × α | x.1⁻¹ * x.2 ∈ U i } := by + rw [uniformity_eq_comap_inv_mul_nhds_one] + exact h.comap _ + +@[to_additive] +theorem Filter.HasBasis.uniformity_of_nhds_one_swapped {ι} {p : ι → Prop} {U : ι → Set α} + (h : (𝓝 (1 : α)).HasBasis p U) : + (𝓤 α).HasBasis p fun i => { x : α × α | x.1 / x.2 ∈ U i } := by + rw [uniformity_eq_comap_nhds_one_swapped] + exact h.comap _ + +@[to_additive] +theorem Filter.HasBasis.uniformity_of_nhds_one_inv_mul_swapped {ι} {p : ι → Prop} {U : ι → Set α} + (h : (𝓝 (1 : α)).HasBasis p U) : + (𝓤 α).HasBasis p fun i => { x : α × α | x.2⁻¹ * x.1 ∈ U i } := by + rw [uniformity_eq_comap_inv_mul_nhds_one_swapped] + exact h.comap _ + +@[to_additive] +theorem uniformContinuous_of_tendsto_one {hom : Type*} [UniformSpace β] [Group β] [UniformGroup β] + [FunLike hom α β] [MonoidHomClass hom α β] {f : hom} (h : Tendsto f (𝓝 1) (𝓝 1)) : + UniformContinuous f := by + have : + ((fun x : β × β => x.2 / x.1) ∘ fun x : α × α => (f x.1, f x.2)) = fun x : α × α => + f (x.2 / x.1) := by ext; simp only [Function.comp_apply, map_div] + rw [UniformContinuous, uniformity_eq_comap_nhds_one α, uniformity_eq_comap_nhds_one β, + tendsto_comap_iff, this] + exact Tendsto.comp h tendsto_comap + +/-- A group homomorphism (a bundled morphism of a type that implements `MonoidHomClass`) between +two uniform groups is uniformly continuous provided that it is continuous at one. See also +`continuous_of_continuousAt_one`. -/ +@[to_additive "An additive group homomorphism (a bundled morphism of a type that implements +`AddMonoidHomClass`) between two uniform additive groups is uniformly continuous provided that it +is continuous at zero. See also `continuous_of_continuousAt_zero`."] +theorem uniformContinuous_of_continuousAt_one {hom : Type*} [UniformSpace β] [Group β] + [UniformGroup β] [FunLike hom α β] [MonoidHomClass hom α β] + (f : hom) (hf : ContinuousAt f 1) : + UniformContinuous f := + uniformContinuous_of_tendsto_one (by simpa using hf.tendsto) + +@[to_additive] +theorem MonoidHom.uniformContinuous_of_continuousAt_one [UniformSpace β] [Group β] [UniformGroup β] + (f : α →* β) (hf : ContinuousAt f 1) : UniformContinuous f := + _root_.uniformContinuous_of_continuousAt_one f hf + +/-- A homomorphism from a uniform group to a discrete uniform group is continuous if and only if +its kernel is open. -/ +@[to_additive "A homomorphism from a uniform additive group to a discrete uniform additive group is +continuous if and only if its kernel is open."] +theorem UniformGroup.uniformContinuous_iff_open_ker {hom : Type*} [UniformSpace β] + [DiscreteTopology β] [Group β] [UniformGroup β] [FunLike hom α β] [MonoidHomClass hom α β] + {f : hom} : + UniformContinuous f ↔ IsOpen ((f : α →* β).ker : Set α) := by + refine ⟨fun hf => ?_, fun hf => ?_⟩ + · apply (isOpen_discrete ({1} : Set β)).preimage hf.continuous + · apply uniformContinuous_of_continuousAt_one + rw [ContinuousAt, nhds_discrete β, map_one, tendsto_pure] + exact hf.mem_nhds (map_one f) + +@[to_additive] +theorem uniformContinuous_monoidHom_of_continuous {hom : Type*} [UniformSpace β] [Group β] + [UniformGroup β] [FunLike hom α β] [MonoidHomClass hom α β] {f : hom} (h : Continuous f) : + UniformContinuous f := + uniformContinuous_of_tendsto_one <| + suffices Tendsto f (𝓝 1) (𝓝 (f 1)) by rwa [map_one] at this + h.tendsto 1 + +end UniformGroup + +section TopologicalGroup + +open Filter + +variable (G : Type*) [Group G] [TopologicalSpace G] [TopologicalGroup G] + +/-- The right uniformity on a topological group (as opposed to the left uniformity). + +Warning: in general the right and left uniformities do not coincide and so one does not obtain a +`UniformGroup` structure. Two important special cases where they _do_ coincide are for +commutative groups (see `comm_topologicalGroup_is_uniform`) and for compact groups (see +`topologicalGroup_is_uniform_of_compactSpace`). -/ +@[to_additive "The right uniformity on a topological additive group (as opposed to the left +uniformity). + +Warning: in general the right and left uniformities do not coincide and so one does not obtain a +`UniformAddGroup` structure. Two important special cases where they _do_ coincide are for +commutative additive groups (see `comm_topologicalAddGroup_is_uniform`) and for compact +additive groups (see `topologicalAddGroup_is_uniform_of_compactSpace`)."] +def TopologicalGroup.toUniformSpace : UniformSpace G where + uniformity := comap (fun p : G × G => p.2 / p.1) (𝓝 1) + symm := + have : Tendsto (fun p : G × G ↦ (p.2 / p.1)⁻¹) (comap (fun p : G × G ↦ p.2 / p.1) (𝓝 1)) + (𝓝 1⁻¹) := tendsto_id.inv.comp tendsto_comap + by simpa [tendsto_comap_iff] + comp := Tendsto.le_comap fun U H ↦ by + rcases exists_nhds_one_split H with ⟨V, V_nhds, V_mul⟩ + refine mem_map.2 (mem_of_superset (mem_lift' <| preimage_mem_comap V_nhds) ?_) + rintro ⟨x, y⟩ ⟨z, hz₁, hz₂⟩ + simpa using V_mul _ hz₂ _ hz₁ + nhds_eq_comap_uniformity _ := by simp only [comap_comap, Function.comp_def, nhds_translation_div] + +attribute [local instance] TopologicalGroup.toUniformSpace + +@[to_additive] +theorem uniformity_eq_comap_nhds_one' : 𝓤 G = comap (fun p : G × G => p.2 / p.1) (𝓝 (1 : G)) := + rfl + +end TopologicalGroup + +section TopologicalCommGroup + +universe u v w x + +open Filter + +variable (G : Type*) [CommGroup G] [TopologicalSpace G] [TopologicalGroup G] + +section + +attribute [local instance] TopologicalGroup.toUniformSpace + +variable {G} + +@[to_additive] +-- Porting note: renamed theorem to conform to naming convention +theorem comm_topologicalGroup_is_uniform : UniformGroup G := by + have : + Tendsto + ((fun p : G × G => p.1 / p.2) ∘ fun p : (G × G) × G × G => (p.1.2 / p.1.1, p.2.2 / p.2.1)) + (comap (fun p : (G × G) × G × G => (p.1.2 / p.1.1, p.2.2 / p.2.1)) ((𝓝 1).prod (𝓝 1))) + (𝓝 (1 / 1)) := + (tendsto_fst.div' tendsto_snd).comp tendsto_comap + constructor + rw [UniformContinuous, uniformity_prod_eq_prod, tendsto_map'_iff, uniformity_eq_comap_nhds_one' G, + tendsto_comap_iff, prod_comap_comap_eq] + simp only [Function.comp_def, div_eq_mul_inv, mul_inv_rev, inv_inv, mul_comm, mul_left_comm] at * + simp only [inv_one, mul_one, ← mul_assoc] at this + simp_rw [← mul_assoc, mul_comm] + assumption + +open Set + +end + +@[to_additive] +theorem UniformGroup.toUniformSpace_eq {G : Type*} [u : UniformSpace G] [Group G] + [UniformGroup G] : TopologicalGroup.toUniformSpace G = u := by + ext : 1 + rw [uniformity_eq_comap_nhds_one' G, uniformity_eq_comap_nhds_one G] + +end TopologicalCommGroup + +open Filter Set Function + +section + +variable {α : Type*} {β : Type*} {hom : Type*} +variable [TopologicalSpace α] [Group α] [TopologicalGroup α] + +-- β is a dense subgroup of α, inclusion is denoted by e +variable [TopologicalSpace β] [Group β] +variable [FunLike hom β α] [MonoidHomClass hom β α] {e : hom} + +@[to_additive] +theorem tendsto_div_comap_self (de : IsDenseInducing e) (x₀ : α) : + Tendsto (fun t : β × β => t.2 / t.1) ((comap fun p : β × β => (e p.1, e p.2)) <| 𝓝 (x₀, x₀)) + (𝓝 1) := by + have comm : ((fun x : α × α => x.2 / x.1) ∘ fun t : β × β => (e t.1, e t.2)) = + e ∘ fun t : β × β => t.2 / t.1 := by + ext t + change e t.2 / e t.1 = e (t.2 / t.1) + rw [← map_div e t.2 t.1] + have lim : Tendsto (fun x : α × α => x.2 / x.1) (𝓝 (x₀, x₀)) (𝓝 (e 1)) := by + simpa using (continuous_div'.comp (@continuous_swap α α _ _)).tendsto (x₀, x₀) + simpa using de.tendsto_comap_nhds_nhds lim comm + +end + +namespace IsDenseInducing + +variable {α : Type*} {β : Type*} {γ : Type*} {δ : Type*} +variable {G : Type*} + +-- β is a dense subgroup of α, inclusion is denoted by e +-- δ is a dense subgroup of γ, inclusion is denoted by f +variable [TopologicalSpace α] [AddCommGroup α] [TopologicalAddGroup α] +variable [TopologicalSpace β] [AddCommGroup β] +variable [TopologicalSpace γ] [AddCommGroup γ] [TopologicalAddGroup γ] +variable [TopologicalSpace δ] [AddCommGroup δ] +variable [UniformSpace G] [AddCommGroup G] +variable {e : β →+ α} (de : IsDenseInducing e) +variable {f : δ →+ γ} (df : IsDenseInducing f) +variable {φ : β →+ δ →+ G} +variable (hφ : Continuous (fun p : β × δ => φ p.1 p.2)) +variable {W' : Set G} (W'_nhd : W' ∈ 𝓝 (0 : G)) +include de hφ + +include W'_nhd in +private theorem extend_Z_bilin_aux (x₀ : α) (y₁ : δ) : ∃ U₂ ∈ comap e (𝓝 x₀), ∀ x ∈ U₂, ∀ x' ∈ U₂, + (fun p : β × δ => φ p.1 p.2) (x' - x, y₁) ∈ W' := by + let Nx := 𝓝 x₀ + let ee := fun u : β × β => (e u.1, e u.2) + have lim1 : Tendsto (fun a : β × β => (a.2 - a.1, y₁)) + (comap e Nx ×ˢ comap e Nx) (𝓝 (0, y₁)) := by + have := Tendsto.prod_mk (tendsto_sub_comap_self de x₀) + (tendsto_const_nhds : Tendsto (fun _ : β × β => y₁) (comap ee <| 𝓝 (x₀, x₀)) (𝓝 y₁)) + rw [nhds_prod_eq, prod_comap_comap_eq, ← nhds_prod_eq] + exact (this : _) + have lim2 : Tendsto (fun p : β × δ => φ p.1 p.2) (𝓝 (0, y₁)) (𝓝 0) := by + simpa using hφ.tendsto (0, y₁) + have lim := lim2.comp lim1 + rw [tendsto_prod_self_iff] at lim + simp_rw [forall_mem_comm] + exact lim W' W'_nhd + +variable [UniformAddGroup G] + +include df W'_nhd in +private theorem extend_Z_bilin_key (x₀ : α) (y₀ : γ) : ∃ U ∈ comap e (𝓝 x₀), ∃ V ∈ comap f (𝓝 y₀), + ∀ x ∈ U, ∀ x' ∈ U, ∀ (y) (_ : y ∈ V) (y') (_ : y' ∈ V), + (fun p : β × δ => φ p.1 p.2) (x', y') - (fun p : β × δ => φ p.1 p.2) (x, y) ∈ W' := by + let ee := fun u : β × β => (e u.1, e u.2) + let ff := fun u : δ × δ => (f u.1, f u.2) + have lim_φ : Filter.Tendsto (fun p : β × δ => φ p.1 p.2) (𝓝 (0, 0)) (𝓝 0) := by + simpa using hφ.tendsto (0, 0) + have lim_φ_sub_sub : + Tendsto (fun p : (β × β) × δ × δ => (fun p : β × δ => φ p.1 p.2) (p.1.2 - p.1.1, p.2.2 - p.2.1)) + ((comap ee <| 𝓝 (x₀, x₀)) ×ˢ (comap ff <| 𝓝 (y₀, y₀))) (𝓝 0) := by + have lim_sub_sub : + Tendsto (fun p : (β × β) × δ × δ => (p.1.2 - p.1.1, p.2.2 - p.2.1)) + (comap ee (𝓝 (x₀, x₀)) ×ˢ comap ff (𝓝 (y₀, y₀))) (𝓝 0 ×ˢ 𝓝 0) := by + have := Filter.prod_mono (tendsto_sub_comap_self de x₀) (tendsto_sub_comap_self df y₀) + rwa [prod_map_map_eq] at this + rw [← nhds_prod_eq] at lim_sub_sub + exact Tendsto.comp lim_φ lim_sub_sub + rcases exists_nhds_zero_quarter W'_nhd with ⟨W, W_nhd, W4⟩ + have : + ∃ U₁ ∈ comap e (𝓝 x₀), ∃ V₁ ∈ comap f (𝓝 y₀), ∀ (x) (_ : x ∈ U₁) (x') (_ : x' ∈ U₁), + ∀ (y) (_ : y ∈ V₁) (y') (_ : y' ∈ V₁), (fun p : β × δ => φ p.1 p.2) (x' - x, y' - y) ∈ W := by + rcases tendsto_prod_iff.1 lim_φ_sub_sub W W_nhd with ⟨U, U_in, V, V_in, H⟩ + rw [nhds_prod_eq, ← prod_comap_comap_eq, mem_prod_same_iff] at U_in V_in + rcases U_in with ⟨U₁, U₁_in, HU₁⟩ + rcases V_in with ⟨V₁, V₁_in, HV₁⟩ + exists U₁, U₁_in, V₁, V₁_in + intro x x_in x' x'_in y y_in y' y'_in + exact H _ _ (HU₁ (mk_mem_prod x_in x'_in)) (HV₁ (mk_mem_prod y_in y'_in)) + rcases this with ⟨U₁, U₁_nhd, V₁, V₁_nhd, H⟩ + obtain ⟨x₁, x₁_in⟩ : U₁.Nonempty := (de.comap_nhds_neBot _).nonempty_of_mem U₁_nhd + obtain ⟨y₁, y₁_in⟩ : V₁.Nonempty := (df.comap_nhds_neBot _).nonempty_of_mem V₁_nhd + have cont_flip : Continuous fun p : δ × β => φ.flip p.1 p.2 := by + show Continuous ((fun p : β × δ => φ p.1 p.2) ∘ Prod.swap) + exact hφ.comp continuous_swap + rcases extend_Z_bilin_aux de hφ W_nhd x₀ y₁ with ⟨U₂, U₂_nhd, HU⟩ + rcases extend_Z_bilin_aux df cont_flip W_nhd y₀ x₁ with ⟨V₂, V₂_nhd, HV⟩ + exists U₁ ∩ U₂, inter_mem U₁_nhd U₂_nhd, V₁ ∩ V₂, inter_mem V₁_nhd V₂_nhd + rintro x ⟨xU₁, xU₂⟩ x' ⟨x'U₁, x'U₂⟩ y ⟨yV₁, yV₂⟩ y' ⟨y'V₁, y'V₂⟩ + have key_formula : φ x' y' - φ x y + = φ (x' - x) y₁ + φ (x' - x) (y' - y₁) + φ x₁ (y' - y) + φ (x - x₁) (y' - y) := by simp; abel + rw [key_formula] + have h₁ := HU x xU₂ x' x'U₂ + have h₂ := H x xU₁ x' x'U₁ y₁ y₁_in y' y'V₁ + have h₃ := HV y yV₂ y' y'V₂ + have h₄ := H x₁ x₁_in x xU₁ y yV₁ y' y'V₁ + exact W4 h₁ h₂ h₃ h₄ + +end IsDenseInducing diff --git a/Mathlib/Topology/Algebra/UniformMulAction.lean b/Mathlib/Topology/Algebra/UniformMulAction.lean index 7fb5a47c14347..e64a29e9f4052 100644 --- a/Mathlib/Topology/Algebra/UniformMulAction.lean +++ b/Mathlib/Topology/Algebra/UniformMulAction.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import Mathlib.Algebra.Module.Opposite -import Mathlib.Topology.Algebra.UniformGroup import Mathlib.Topology.UniformSpace.Completion +import Mathlib.Topology.Algebra.UniformGroup.Defs /-! # Multiplicative action on the completion of a uniform space diff --git a/Mathlib/Topology/Algebra/UniformRing.lean b/Mathlib/Topology/Algebra/UniformRing.lean index 7ad0d76af8d94..3c361e58e78d7 100644 --- a/Mathlib/Topology/Algebra/UniformRing.lean +++ b/Mathlib/Topology/Algebra/UniformRing.lean @@ -7,6 +7,7 @@ import Mathlib.Algebra.Algebra.Defs import Mathlib.Logic.Equiv.TransferInstance import Mathlib.Topology.Algebra.GroupCompletion import Mathlib.Topology.Algebra.Ring.Ideal +import Mathlib.Topology.Algebra.UniformGroup.Basic /-! # Completion of topological rings: diff --git a/Mathlib/Topology/ContinuousMap/Algebra.lean b/Mathlib/Topology/ContinuousMap/Algebra.lean index c36b582bb8332..c8e869fee7be1 100644 --- a/Mathlib/Topology/ContinuousMap/Algebra.lean +++ b/Mathlib/Topology/ContinuousMap/Algebra.lean @@ -13,7 +13,6 @@ import Mathlib.Topology.Algebra.Module.Basic import Mathlib.Topology.Algebra.InfiniteSum.Basic import Mathlib.Topology.Algebra.Ring.Basic import Mathlib.Topology.Algebra.Star -import Mathlib.Topology.Algebra.UniformGroup import Mathlib.Topology.ContinuousMap.Ordered import Mathlib.Topology.UniformSpace.CompactConvergence diff --git a/Mathlib/Topology/Instances/Real.lean b/Mathlib/Topology/Instances/Real.lean index c41f014575333..1d2124765c1cf 100644 --- a/Mathlib/Topology/Instances/Real.lean +++ b/Mathlib/Topology/Instances/Real.lean @@ -14,6 +14,7 @@ import Mathlib.Topology.Algebra.UniformMulAction import Mathlib.Topology.Instances.Int import Mathlib.Topology.Metrizable.Basic import Mathlib.Topology.Order.Bornology +import Mathlib.Topology.Algebra.UniformGroup.Basic /-! # Topological properties of ℝ diff --git a/Mathlib/Topology/UniformSpace/Matrix.lean b/Mathlib/Topology/UniformSpace/Matrix.lean index 4afd736c939f3..cff2509e52dc9 100644 --- a/Mathlib/Topology/UniformSpace/Matrix.lean +++ b/Mathlib/Topology/UniformSpace/Matrix.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Eric Wieser. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Eric Wieser, Heather Macbeth -/ -import Mathlib.Topology.Algebra.UniformGroup +import Mathlib.Topology.Algebra.UniformGroup.Basic import Mathlib.Topology.UniformSpace.Pi import Mathlib.Data.Matrix.Basic From deda35aa3eb0ef0a1b8f3c21f17169bf4e785cdf Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 28 Oct 2024 06:12:41 +0000 Subject: [PATCH 504/505] chore: split Data.Set.Basic (#18317) Delay imports of `Data.Order.SymmDiff`, as many things downstream don't need it. --- Mathlib.lean | 1 + Mathlib/Algebra/Group/Center.lean | 2 +- Mathlib/Algebra/Group/Indicator.lean | 1 + Mathlib/Data/Finset/Basic.lean | 1 + Mathlib/Data/PEquiv.lean | 3 +- Mathlib/Data/PNat/Basic.lean | 1 - Mathlib/Data/Set/Basic.lean | 40 +------------------ Mathlib/Data/Set/Enumerate.lean | 2 +- Mathlib/Data/Set/Image.lean | 1 + Mathlib/Data/Set/SymmDiff.lean | 55 +++++++++++++++++++++++++++ Mathlib/Data/SetLike/Basic.lean | 2 +- Mathlib/Order/Filter/Defs.lean | 2 +- Mathlib/Order/Interval/Set/Basic.lean | 1 + Mathlib/Topology/Connected/Basic.lean | 2 +- scripts/noshake.json | 6 ++- 15 files changed, 73 insertions(+), 47 deletions(-) create mode 100644 Mathlib/Data/Set/SymmDiff.lean diff --git a/Mathlib.lean b/Mathlib.lean index 444f67315b74f..235cfb076161b 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2688,6 +2688,7 @@ import Mathlib.Data.Set.Sigma import Mathlib.Data.Set.Subset import Mathlib.Data.Set.Subsingleton import Mathlib.Data.Set.Sups +import Mathlib.Data.Set.SymmDiff import Mathlib.Data.Set.UnionLift import Mathlib.Data.SetLike.Basic import Mathlib.Data.SetLike.Fintype diff --git a/Mathlib/Algebra/Group/Center.lean b/Mathlib/Algebra/Group/Center.lean index 786c03ef1f1db..d5fd79ac44d8c 100644 --- a/Mathlib/Algebra/Group/Center.lean +++ b/Mathlib/Algebra/Group/Center.lean @@ -5,8 +5,8 @@ Authors: Eric Wieser, Jireh Loreaux -/ import Mathlib.Algebra.Group.Commute.Units import Mathlib.Algebra.Group.Invertible.Basic -import Mathlib.Data.Set.Basic import Mathlib.Logic.Basic +import Mathlib.Data.Set.Basic /-! # Centers of magmas and semigroups diff --git a/Mathlib/Algebra/Group/Indicator.lean b/Mathlib/Algebra/Group/Indicator.lean index 38b9e8ce30021..b30e9c0da0970 100644 --- a/Mathlib/Algebra/Group/Indicator.lean +++ b/Mathlib/Algebra/Group/Indicator.lean @@ -5,6 +5,7 @@ Authors: Zhouhang Zhou -/ import Mathlib.Algebra.Group.Pi.Lemmas import Mathlib.Algebra.Group.Support +import Mathlib.Data.Set.SymmDiff /-! # Indicator function diff --git a/Mathlib/Data/Finset/Basic.lean b/Mathlib/Data/Finset/Basic.lean index e3a1af524b590..20b17a33c6b37 100644 --- a/Mathlib/Data/Finset/Basic.lean +++ b/Mathlib/Data/Finset/Basic.lean @@ -8,6 +8,7 @@ import Mathlib.Data.Multiset.FinsetOps import Mathlib.Logic.Equiv.Set import Mathlib.Order.Directed import Mathlib.Order.Interval.Set.Basic +import Mathlib.Data.Set.SymmDiff /-! # Finite sets diff --git a/Mathlib/Data/PEquiv.lean b/Mathlib/Data/PEquiv.lean index 3228b351dff8f..28b8f8af64e9b 100644 --- a/Mathlib/Data/PEquiv.lean +++ b/Mathlib/Data/PEquiv.lean @@ -4,8 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes -/ import Mathlib.Data.Option.Basic -import Mathlib.Data.Set.Basic import Batteries.Tactic.Congr +import Mathlib.Data.Set.Basic +import Mathlib.Tactic.Contrapose /-! diff --git a/Mathlib/Data/PNat/Basic.lean b/Mathlib/Data/PNat/Basic.lean index 05f07c00bfc81..f81e93cd63db2 100644 --- a/Mathlib/Data/PNat/Basic.lean +++ b/Mathlib/Data/PNat/Basic.lean @@ -5,7 +5,6 @@ Authors: Mario Carneiro, Ralf Stephan, Neil Strickland, Ruben Van de Velde -/ import Mathlib.Data.PNat.Equiv import Mathlib.Algebra.Order.Ring.Nat -import Mathlib.Data.Set.Basic import Mathlib.Algebra.GroupWithZero.Divisibility import Mathlib.Algebra.Order.Positive.Ring import Mathlib.Order.Hom.Basic diff --git a/Mathlib/Data/Set/Basic.lean b/Mathlib/Data/Set/Basic.lean index f755b5a1cb2a8..91f306ee2b92d 100644 --- a/Mathlib/Data/Set/Basic.lean +++ b/Mathlib/Data/Set/Basic.lean @@ -6,10 +6,11 @@ Authors: Jeremy Avigad, Leonardo de Moura import Mathlib.Algebra.Group.ZeroOne import Mathlib.Data.Set.Operations import Mathlib.Order.Basic -import Mathlib.Order.SymmDiff +import Mathlib.Order.BooleanAlgebra import Mathlib.Tactic.Tauto import Mathlib.Tactic.ByContra import Mathlib.Util.Delaborators +import Mathlib.Tactic.Lift /-! # Basic properties of sets @@ -1695,43 +1696,6 @@ theorem Nonempty.subset_pair_iff_eq (hs : s.Nonempty) : s ⊆ {a, b} ↔ s = {a} ∨ s = {b} ∨ s = {a, b} := by rw [Set.subset_pair_iff_eq, or_iff_right]; exact hs.ne_empty -/-! ### Symmetric difference -/ - -section - -open scoped symmDiff - -theorem mem_symmDiff : a ∈ s ∆ t ↔ a ∈ s ∧ a ∉ t ∨ a ∈ t ∧ a ∉ s := - Iff.rfl - -protected theorem symmDiff_def (s t : Set α) : s ∆ t = s \ t ∪ t \ s := - rfl - -theorem symmDiff_subset_union : s ∆ t ⊆ s ∪ t := - @symmDiff_le_sup (Set α) _ _ _ - -@[simp] -theorem symmDiff_eq_empty : s ∆ t = ∅ ↔ s = t := - symmDiff_eq_bot - -@[simp] -theorem symmDiff_nonempty : (s ∆ t).Nonempty ↔ s ≠ t := - nonempty_iff_ne_empty.trans symmDiff_eq_empty.not - -theorem inter_symmDiff_distrib_left (s t u : Set α) : s ∩ t ∆ u = (s ∩ t) ∆ (s ∩ u) := - inf_symmDiff_distrib_left _ _ _ - -theorem inter_symmDiff_distrib_right (s t u : Set α) : s ∆ t ∩ u = (s ∩ u) ∆ (t ∩ u) := - inf_symmDiff_distrib_right _ _ _ - -theorem subset_symmDiff_union_symmDiff_left (h : Disjoint s t) : u ⊆ s ∆ u ∪ t ∆ u := - h.le_symmDiff_sup_symmDiff_left - -theorem subset_symmDiff_union_symmDiff_right (h : Disjoint t u) : s ⊆ s ∆ t ∪ s ∆ u := - h.le_symmDiff_sup_symmDiff_right - -end - /-! ### Powerset -/ theorem mem_powerset {x s : Set α} (h : x ⊆ s) : x ∈ 𝒫 s := @h diff --git a/Mathlib/Data/Set/Enumerate.lean b/Mathlib/Data/Set/Enumerate.lean index f5c71a78aa47e..908974989a9e3 100644 --- a/Mathlib/Data/Set/Enumerate.lean +++ b/Mathlib/Data/Set/Enumerate.lean @@ -5,8 +5,8 @@ Authors: Johannes Hölzl -/ import Mathlib.Algebra.Group.Basic import Mathlib.Algebra.Group.Nat -import Mathlib.Data.Set.Basic import Mathlib.Tactic.Common +import Mathlib.Data.Set.Basic /-! # Set enumeration diff --git a/Mathlib/Data/Set/Image.lean b/Mathlib/Data/Set/Image.lean index 1da452c8964d3..af9427b7d8259 100644 --- a/Mathlib/Data/Set/Image.lean +++ b/Mathlib/Data/Set/Image.lean @@ -8,6 +8,7 @@ import Mathlib.Tactic.Use import Batteries.Tactic.Congr import Mathlib.Order.TypeTags import Mathlib.Data.Option.Basic +import Mathlib.Data.Set.SymmDiff /-! # Images and preimages of sets diff --git a/Mathlib/Data/Set/SymmDiff.lean b/Mathlib/Data/Set/SymmDiff.lean new file mode 100644 index 0000000000000..7914ca1b3143d --- /dev/null +++ b/Mathlib/Data/Set/SymmDiff.lean @@ -0,0 +1,55 @@ +/- +Copyright (c) 2014 Jeremy Avigad. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jeremy Avigad, Leonardo de Moura +-/ +import Mathlib.Algebra.Group.ZeroOne +import Mathlib.Data.Set.Operations +import Mathlib.Order.Basic +import Mathlib.Order.BooleanAlgebra +import Mathlib.Logic.Equiv.Basic +import Mathlib.Tactic.Tauto +import Mathlib.Tactic.ByContra +import Mathlib.Util.Delaborators +import Mathlib.Order.SymmDiff +import Mathlib.Data.Set.Basic + +/-! # Symmetric differences of sets -/ + +namespace Set + +universe u v +variable {α : Type u} {β : Type v} {a b : α} {s s₁ s₂ t t₁ t₂ u : Set α} + +open scoped symmDiff + +theorem mem_symmDiff : a ∈ s ∆ t ↔ a ∈ s ∧ a ∉ t ∨ a ∈ t ∧ a ∉ s := + Iff.rfl + +protected theorem symmDiff_def (s t : Set α) : s ∆ t = s \ t ∪ t \ s := + rfl + +theorem symmDiff_subset_union : s ∆ t ⊆ s ∪ t := + @symmDiff_le_sup (Set α) _ _ _ + +@[simp] +theorem symmDiff_eq_empty : s ∆ t = ∅ ↔ s = t := + symmDiff_eq_bot + +@[simp] +theorem symmDiff_nonempty : (s ∆ t).Nonempty ↔ s ≠ t := + nonempty_iff_ne_empty.trans symmDiff_eq_empty.not + +theorem inter_symmDiff_distrib_left (s t u : Set α) : s ∩ t ∆ u = (s ∩ t) ∆ (s ∩ u) := + inf_symmDiff_distrib_left _ _ _ + +theorem inter_symmDiff_distrib_right (s t u : Set α) : s ∆ t ∩ u = (s ∩ u) ∆ (t ∩ u) := + inf_symmDiff_distrib_right _ _ _ + +theorem subset_symmDiff_union_symmDiff_left (h : Disjoint s t) : u ⊆ s ∆ u ∪ t ∆ u := + h.le_symmDiff_sup_symmDiff_left + +theorem subset_symmDiff_union_symmDiff_right (h : Disjoint t u) : s ⊆ s ∆ t ∪ s ∆ u := + h.le_symmDiff_sup_symmDiff_right + +end Set diff --git a/Mathlib/Data/SetLike/Basic.lean b/Mathlib/Data/SetLike/Basic.lean index 2a07db5d06280..021b2c1a7978b 100644 --- a/Mathlib/Data/SetLike/Basic.lean +++ b/Mathlib/Data/SetLike/Basic.lean @@ -3,9 +3,9 @@ Copyright (c) 2021 Eric Wieser. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Eric Wieser -/ -import Mathlib.Data.Set.Basic import Mathlib.Tactic.Monotonicity.Attr import Mathlib.Tactic.SetLike +import Mathlib.Data.Set.Basic /-! # Typeclass for types with a set-like extensionality property diff --git a/Mathlib/Order/Filter/Defs.lean b/Mathlib/Order/Filter/Defs.lean index 4b9a6cb4f9a6b..a91d2c91457f3 100644 --- a/Mathlib/Order/Filter/Defs.lean +++ b/Mathlib/Order/Filter/Defs.lean @@ -3,9 +3,9 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Jeremy Avigad -/ -import Mathlib.Data.Set.Basic import Mathlib.Order.SetNotation import Mathlib.Order.Bounds.Defs +import Mathlib.Data.Set.Basic /-! # Definitions about filters diff --git a/Mathlib/Order/Interval/Set/Basic.lean b/Mathlib/Order/Interval/Set/Basic.lean index 8569f24b3fdc5..0c9a2428277fb 100644 --- a/Mathlib/Order/Interval/Set/Basic.lean +++ b/Mathlib/Order/Interval/Set/Basic.lean @@ -6,6 +6,7 @@ Authors: Johannes Hölzl, Mario Carneiro, Patrick Massot, Yury Kudryashov, Rémy import Mathlib.Order.MinMax import Mathlib.Data.Set.Subsingleton import Mathlib.Tactic.Says +import Mathlib.Tactic.Contrapose /-! # Intervals diff --git a/Mathlib/Topology/Connected/Basic.lean b/Mathlib/Topology/Connected/Basic.lean index 9cb32af224c6e..e09d49be8fb44 100644 --- a/Mathlib/Topology/Connected/Basic.lean +++ b/Mathlib/Topology/Connected/Basic.lean @@ -3,7 +3,7 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro, Yury Kudryashov -/ -import Mathlib.Data.Set.Basic +import Mathlib.Data.Set.SymmDiff import Mathlib.Order.SuccPred.Relation import Mathlib.Topology.Irreducible diff --git a/scripts/noshake.json b/scripts/noshake.json index e6f8ba397178e..b1e321c746453 100644 --- a/scripts/noshake.json +++ b/scripts/noshake.json @@ -219,7 +219,8 @@ "Mathlib.Tactic.ProxyType": ["Mathlib.Logic.Equiv.Defs"], "Mathlib.Tactic.ProdAssoc": ["Mathlib.Logic.Equiv.Defs"], "Mathlib.Tactic.Positivity.Finset": - ["Mathlib.Algebra.Order.BigOperators.Group.Finset", "Mathlib.Data.Finset.Density"], + ["Mathlib.Algebra.Order.BigOperators.Group.Finset", + "Mathlib.Data.Finset.Density"], "Mathlib.Tactic.Positivity.Basic": ["Mathlib.Algebra.GroupPower.Order", "Mathlib.Algebra.Order.Group.PosPart", @@ -350,7 +351,8 @@ "Mathlib.Deprecated.MinMax": ["Mathlib.Order.MinMax"], "Mathlib.Deprecated.ByteArray": ["Batteries.Data.ByteSubarray"], "Mathlib.Data.Vector.Basic": ["Mathlib.Control.Applicative"], - "Mathlib.Data.Set.Image": ["Batteries.Tactic.Congr"], + "Mathlib.Data.Set.Image": + ["Batteries.Tactic.Congr", "Mathlib.Data.Set.SymmDiff"], "Mathlib.Data.Seq.Parallel": ["Mathlib.Init.Data.Prod"], "Mathlib.Data.PEquiv": ["Batteries.Tactic.Congr"], "Mathlib.Data.Ordmap.Ordnode": ["Mathlib.Data.Option.Basic"], From 9fc60c1ad982f142d6d0671bef8f4cda3b1590cd Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 28 Oct 2024 07:22:41 +0000 Subject: [PATCH 505/505] feat: `ZMod`-module lemmas (#17693) From PFR --- Mathlib/Data/ZMod/Basic.lean | 61 ++++++++++++++++++++++++++++++++++-- 1 file changed, 58 insertions(+), 3 deletions(-) diff --git a/Mathlib/Data/ZMod/Basic.lean b/Mathlib/Data/ZMod/Basic.lean index 0200eb971668e..3edc84886e2cc 100644 --- a/Mathlib/Data/ZMod/Basic.lean +++ b/Mathlib/Data/ZMod/Basic.lean @@ -35,7 +35,7 @@ This is a ring hom if the ring has characteristic dividing `n` assert_not_exists Submodule -open Function +open Function ZMod namespace ZMod @@ -1431,9 +1431,26 @@ end lift end ZMod +/-! +### Groups of bounded torsion + +For `G` a group and `n` a natural number, `G` having torsion dividing `n` +(`∀ x : G, n • x = 0`) can be derived from `Module R G` where `R` has characteristic dividing `n`. + +It is however painful to have the API for such groups `G` stated in this generality, as `R` does not +appear anywhere in the lemmas' return type. Instead of writing the API in terms of a general `R`, we +therefore specialise to the canonical ring of order `n`, namely `ZMod n`. + +This spelling `Module (ZMod n) G` has the extra advantage of providing the canonical action by +`ZMod n`. It is however Type-valued, so we might want to acquire a Prop-valued version in the +future. +-/ + section Module -variable {S G : Type*} [AddCommGroup G] {n : ℕ} [Module (ZMod n) G] [SetLike S G] - [AddSubgroupClass S G] {K : S} {x : G} +variable {n : ℕ} {S G : Type*} [AddCommGroup G] [SetLike S G] [AddSubgroupClass S G] {K : S} {x : G} + +section general +variable [Module (ZMod n) G] {x : G} lemma zmod_smul_mem (hx : x ∈ K) : ∀ a : ZMod n, a • x ∈ K := by simpa [ZMod.forall, Int.cast_smul_eq_zsmul] using zsmul_mem hx @@ -1452,6 +1469,44 @@ instance instZModModule : Module (ZMod n) K := Subtype.coe_injective.module _ (AddSubmonoidClass.subtype K) coe_zmod_smul end AddSubgroupClass + +variable (n) + +lemma ZModModule.char_nsmul_eq_zero (x : G) : n • x = 0 := by + simp [← Nat.cast_smul_eq_nsmul (ZMod n)] + +variable (G) in +lemma ZModModule.char_ne_one [Nontrivial G] : n ≠ 1 := by + rintro rfl + obtain ⟨x, hx⟩ := exists_ne (0 : G) + exact hx <| by simpa using char_nsmul_eq_zero 1 x + +variable (G) in +lemma ZModModule.two_le_char [NeZero n] [Nontrivial G] : 2 ≤ n := by + have := NeZero.ne n + have := char_ne_one n G + omega + +lemma ZModModule.periodicPts_add_left [NeZero n] (x : G) : periodicPts (x + ·) = .univ := + Set.eq_univ_of_forall fun y ↦ ⟨n, NeZero.pos n, by + simpa [char_nsmul_eq_zero, IsPeriodicPt] using isFixedPt_id _⟩ + +end general + +section two +variable [Module (ZMod 2) G] + +lemma ZModModule.add_self (x : G) : x + x = 0 := by + simpa [two_nsmul] using char_nsmul_eq_zero 2 x + +lemma ZModModule.neg_eq_self (x : G) : -x = x := by simp [add_self, eq_comm, ← sub_eq_zero] + +lemma ZModModule.sub_eq_add (x y : G) : x - y = x + y := by simp [neg_eq_self, sub_eq_add_neg] + +lemma ZModModule.add_add_add_cancel (x y z : G) : (x + y) + (y + z) = x + z := by + simpa [sub_eq_add] using sub_add_sub_cancel x y z + +end two end Module section AddGroup